From 42c2007a0722c5560f92874f40ce5b4edb1bc389 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 15 Nov 2024 22:50:29 +0000 Subject: [PATCH 001/829] chore: do not use numeric projections for isLittleO (#19114) This means the projection can be swapped for a regular function later without breaking these proofs. --- .../Analysis/Calculus/FDeriv/RestrictScalars.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Analysis/Calculus/FDeriv/RestrictScalars.lean b/Mathlib/Analysis/Calculus/FDeriv/RestrictScalars.lean index 18616c9afda61..7936b7024558c 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/RestrictScalars.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/RestrictScalars.lean @@ -45,21 +45,21 @@ variable {f : E → F} {f' : E →L[𝕜'] F} {s : Set E} {x : E} @[fun_prop] theorem HasStrictFDerivAt.restrictScalars (h : HasStrictFDerivAt f f' x) : HasStrictFDerivAt f (f'.restrictScalars 𝕜) x := - .of_isLittleO h.1 + .of_isLittleO h.isLittleO theorem HasFDerivAtFilter.restrictScalars {L} (h : HasFDerivAtFilter f f' x L) : HasFDerivAtFilter f (f'.restrictScalars 𝕜) x L := - .of_isLittleO h.1 + .of_isLittleO h.isLittleO @[fun_prop] theorem HasFDerivAt.restrictScalars (h : HasFDerivAt f f' x) : HasFDerivAt f (f'.restrictScalars 𝕜) x := - .of_isLittleO h.1 + .of_isLittleO h.isLittleO @[fun_prop] theorem HasFDerivWithinAt.restrictScalars (h : HasFDerivWithinAt f f' s x) : HasFDerivWithinAt f (f'.restrictScalars 𝕜) s x := - .of_isLittleO h.1 + .of_isLittleO h.isLittleO @[fun_prop] theorem DifferentiableAt.restrictScalars (h : DifferentiableAt 𝕜' f x) : DifferentiableAt 𝕜 f x := @@ -82,13 +82,13 @@ theorem Differentiable.restrictScalars (h : Differentiable 𝕜' f) : Differenti theorem HasFDerivWithinAt.of_restrictScalars {g' : E →L[𝕜] F} (h : HasFDerivWithinAt f g' s x) (H : f'.restrictScalars 𝕜 = g') : HasFDerivWithinAt f f' s x := by rw [← H] at h - exact .of_isLittleO h.1 + exact .of_isLittleO h.isLittleO @[fun_prop] theorem hasFDerivAt_of_restrictScalars {g' : E →L[𝕜] F} (h : HasFDerivAt f g' x) (H : f'.restrictScalars 𝕜 = g') : HasFDerivAt f f' x := by rw [← H] at h - exact .of_isLittleO h.1 + exact .of_isLittleO h.isLittleO theorem DifferentiableAt.fderiv_restrictScalars (h : DifferentiableAt 𝕜' f x) : fderiv 𝕜 f x = (fderiv 𝕜' f x).restrictScalars 𝕜 := From 39806f723cd767cd001ff108fe2b8959085c3936 Mon Sep 17 00:00:00 2001 From: damiano Date: Fri, 15 Nov 2024 23:03:09 +0000 Subject: [PATCH 002/829] fix(CI): allow trailing spaces in `maintainer merge/delegate` (#19113) Before #18867, `maintainer merge` and `maintainer delegate` checked that a comment had a line beginning with `maintainer merge` or `maintainer delegate` to start the action. #18867 made it (essentially) stricter, only allowing `^maintainer *merge$` or `maintainer *delegate$`. This PR allows trailing spaces, but not arbitrary text following `maintainer xxx`. --- .github/workflows/maintainer_merge.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.github/workflows/maintainer_merge.yml b/.github/workflows/maintainer_merge.yml index 4b254a5aa3a4f..be9e7ee77facb 100644 --- a/.github/workflows/maintainer_merge.yml +++ b/.github/workflows/maintainer_merge.yml @@ -39,7 +39,7 @@ jobs: printf '%s' "${COMMENT}" | hexdump -cC printf 'Comment:"%s"\n' "${COMMENT}" m_or_d="$(printf '%s' "${COMMENT}" | - sed -n 's=^maintainer *\(merge\|delegate\)$=\1=p' | head -1)" + sed -n 's=^maintainer *\(merge\|delegate\) *$=\1=p' | head -1)" printf $'"maintainer delegate" or "maintainer merge" found? \'%s\'\n' "${m_or_d}" From 22b5743395628f2dc74e4acdbbe3856b03053722 Mon Sep 17 00:00:00 2001 From: JonBannon Date: Fri, 15 Nov 2024 23:56:05 +0000 Subject: [PATCH 003/829] feat(Algebra.Order.Star.Basic): add `aesop` lemma `IsSelfAdjoint.mul_self_nonneg`. (#19033) Includes the simple fact that in a `StarOrderedRing`, selfadjoint elements have nonnegative squares. Co-authored-by: JonBannon <59937998+JonBannon@users.noreply.github.com> --- Mathlib/Algebra/Order/Star/Basic.lean | 8 ++++++++ .../ContinuousFunctionalCalculus/Instances.lean | 4 ++-- 2 files changed, 10 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Order/Star/Basic.lean b/Mathlib/Algebra/Order/Star/Basic.lean index df48d4f580a37..0297489633917 100644 --- a/Mathlib/Algebra/Order/Star/Basic.lean +++ b/Mathlib/Algebra/Order/Star/Basic.lean @@ -168,6 +168,10 @@ theorem star_mul_self_nonneg (r : R) : 0 ≤ star r * r := theorem mul_star_self_nonneg (r : R) : 0 ≤ r * star r := by simpa only [star_star] using star_mul_self_nonneg (star r) +@[aesop safe apply (rule_sets := [CStarAlgebra])] +theorem IsSelfAdjoint.mul_self_nonneg {a : R} (ha : IsSelfAdjoint a) : 0 ≤ a * a := by + simpa [ha.star_eq] using star_mul_self_nonneg a + @[aesop safe apply] theorem conjugate_nonneg {a : R} (ha : 0 ≤ a) (c : R) : 0 ≤ star c * a * c := by rw [StarOrderedRing.nonneg_iff] at ha @@ -297,6 +301,10 @@ lemma one_lt_star_iff {x : R} : 1 < star x ↔ 1 < x := by lemma star_lt_one_iff {x : R} : star x < 1 ↔ x < 1 := by simpa using star_lt_star_iff (x := x) (y := 1) +@[aesop safe apply (rule_sets := [CStarAlgebra])] +theorem IsSelfAdjoint.sq_nonneg {a : R} (ha : IsSelfAdjoint a) : 0 ≤ a ^ 2 := by + simp [sq, ha.mul_self_nonneg] + end Semiring section StarModule diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean index 76c6e66606ed2..b50bedc7fc13c 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean @@ -409,11 +409,11 @@ lemma SpectrumRestricts.nnreal_add {a b : A} (ha₁ : IsSelfAdjoint a) gcongr all_goals rw [← SpectrumRestricts.nnreal_iff_nnnorm] <;> first | rfl | assumption -lemma IsSelfAdjoint.sq_spectrumRestricts {a : A} (ha : IsSelfAdjoint a) : +protected lemma IsSelfAdjoint.sq_spectrumRestricts {a : A} (ha : IsSelfAdjoint a) : SpectrumRestricts (a ^ 2) ContinuousMap.realToNNReal := by rw [SpectrumRestricts.nnreal_iff, ← cfc_id (R := ℝ) a, ← cfc_pow .., cfc_map_spectrum ..] rintro - ⟨x, -, rfl⟩ - exact sq_nonneg x + exact _root_.sq_nonneg (id x) open ComplexStarModule From b6d78d3f5c429d4b1ea31530808fd4ba5b3bb2b8 Mon Sep 17 00:00:00 2001 From: damiano Date: Sat, 16 Nov 2024 00:24:40 +0000 Subject: [PATCH 004/829] chore: remove unused variables (#19101) #17715 --- Mathlib/Algebra/Group/Pointwise/Set/Basic.lean | 2 +- Mathlib/Algebra/Group/Subsemigroup/Basic.lean | 3 +-- Mathlib/Algebra/Module/Equiv/Basic.lean | 10 +++------- Mathlib/Algebra/Module/LinearMap/End.lean | 6 +++--- Mathlib/Algebra/Module/Opposite.lean | 2 +- Mathlib/Algebra/Order/Ring/Basic.lean | 2 +- Mathlib/Algebra/Star/StarRingHom.lean | 4 ++-- Mathlib/Analysis/Analytic/Within.lean | 5 ++--- Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean | 8 +++----- Mathlib/Analysis/Calculus/FDeriv/Comp.lean | 7 +------ Mathlib/Analysis/Calculus/FDeriv/Equiv.lean | 7 +------ Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean | 6 +++--- Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean | 10 ++++------ Mathlib/Data/Finset/Insert.lean | 9 ++++----- Mathlib/Data/Finset/Lattice/Basic.lean | 4 ++-- Mathlib/Data/Matroid/Closure.lean | 2 +- Mathlib/Data/Multiset/Basic.lean | 2 +- Mathlib/Data/Set/Functor.lean | 2 +- Mathlib/Data/Set/Sigma.lean | 2 +- Mathlib/Deprecated/Submonoid.lean | 3 +-- Mathlib/ModelTheory/DirectLimit.lean | 2 +- Mathlib/NumberTheory/ModularForms/SlashActions.lean | 2 +- Mathlib/Order/ConditionallyCompleteLattice/Basic.lean | 6 +++--- .../Order/ConditionallyCompleteLattice/Indexed.lean | 6 +++--- Mathlib/Order/Interval/Finset/Box.lean | 2 +- Mathlib/Order/Interval/Set/UnorderedInterval.lean | 2 +- Mathlib/RingTheory/Derivation/Lie.lean | 2 +- Mathlib/Topology/Separation/GDelta.lean | 4 ++-- 28 files changed, 50 insertions(+), 72 deletions(-) diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean index cf1abd5bdefec..187539c08b148 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean @@ -855,7 +855,7 @@ lemma smul_set_iInter₂_subset (a : α) (t : ∀ i, κ i → Set β) : end SMulSet -variable {s s₁ s₂ : Set α} {t t₁ t₂ : Set β} {a : α} {b : β} +variable {s : Set α} {t : Set β} {a : α} {b : β} @[to_additive] lemma range_smul_range {ι κ : Type*} [SMul α β] (b : ι → α) (c : κ → β) : diff --git a/Mathlib/Algebra/Group/Subsemigroup/Basic.lean b/Mathlib/Algebra/Group/Subsemigroup/Basic.lean index 31ab296bcc1c3..72a594ef2bd7a 100644 --- a/Mathlib/Algebra/Group/Subsemigroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subsemigroup/Basic.lean @@ -43,12 +43,11 @@ subsemigroup, subsemigroups assert_not_exists MonoidWithZero -- Only needed for notation -variable {M : Type*} {N : Type*} {A : Type*} +variable {M : Type*} {N : Type*} section NonAssoc variable [Mul M] {s : Set M} -variable [Add A] {t : Set A} namespace Subsemigroup diff --git a/Mathlib/Algebra/Module/Equiv/Basic.lean b/Mathlib/Algebra/Module/Equiv/Basic.lean index b5a356bcc0795..67ae9340c5ef4 100644 --- a/Mathlib/Algebra/Module/Equiv/Basic.lean +++ b/Mathlib/Algebra/Module/Equiv/Basic.lean @@ -18,11 +18,8 @@ import Mathlib.Algebra.Module.Pi open Function -universe u u' v w x y z - -variable {R : Type*} {R₁ : Type*} {R₂ : Type*} {R₃ : Type*} -variable {k : Type*} {K : Type*} {S : Type*} {M : Type*} {M₁ : Type*} {M₂ : Type*} {M₃ : Type*} -variable {N₁ : Type*} {N₂ : Type*} {N₃ : Type*} {N₄ : Type*} {ι : Type*} +variable {R : Type*} {R₂ : Type*} +variable {K : Type*} {S : Type*} {M : Type*} {M₁ : Type*} {M₂ : Type*} {M₃ : Type*} section AddCommMonoid @@ -593,8 +590,7 @@ end CommSemiring section Field -variable [Field K] [AddCommGroup M] [AddCommGroup M₂] [AddCommGroup M₃] -variable [Module K M] [Module K M₂] [Module K M₃] +variable [Field K] [AddCommGroup M] [Module K M] variable (K) (M) open LinearMap diff --git a/Mathlib/Algebra/Module/LinearMap/End.lean b/Mathlib/Algebra/Module/LinearMap/End.lean index f2c1f9a46d3d6..5f1a7fe6fd3d0 100644 --- a/Mathlib/Algebra/Module/LinearMap/End.lean +++ b/Mathlib/Algebra/Module/LinearMap/End.lean @@ -30,7 +30,7 @@ universe u v abbrev Module.End (R : Type u) (M : Type v) [Semiring R] [AddCommMonoid M] [Module R M] := M →ₗ[R] M -variable {R R₁ R₂ S M M₁ M₂ M₃ N N₁ N₂ : Type*} +variable {R R₂ S M M₁ M₂ M₃ N₁ : Type*} namespace LinearMap @@ -301,7 +301,7 @@ section AddCommMonoid section SMulRight -variable [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₁] [Module R M] [Module R M₁] +variable [Semiring R] [AddCommMonoid M] [AddCommMonoid M₁] [Module R M] [Module R M₁] variable [Semiring S] [Module R S] [Module S M] [IsScalarTower R S M] /-- When `f` is an `R`-linear map taking values in `S`, then `fun b ↦ f b • x` is an `R`-linear @@ -363,7 +363,7 @@ section CommSemiring variable [CommSemiring R] [AddCommMonoid M] [AddCommMonoid M₂] [AddCommMonoid M₃] variable [Module R M] [Module R M₂] [Module R M₃] -variable (f g : M →ₗ[R] M₂) +variable (f : M →ₗ[R] M₂) /-- Composition by `f : M₂ → M₃` is a linear map from the space of linear maps `M → M₂` to the space of linear maps `M → M₃`. -/ diff --git a/Mathlib/Algebra/Module/Opposite.lean b/Mathlib/Algebra/Module/Opposite.lean index 5994afcd97b8f..a395a3941b3bc 100644 --- a/Mathlib/Algebra/Module/Opposite.lean +++ b/Mathlib/Algebra/Module/Opposite.lean @@ -18,7 +18,7 @@ assert_not_exists LinearMap section -variable {R S M : Type*} [Semiring R] [Semiring S] [AddCommMonoid M] [Module S M] +variable {R M : Type*} [Semiring R] [AddCommMonoid M] -- see Note [lower instance priority] /-- Like `Semiring.toModule`, but multiplies on the right. -/ diff --git a/Mathlib/Algebra/Order/Ring/Basic.lean b/Mathlib/Algebra/Order/Ring/Basic.lean index 6136beacc1690..542142b07165f 100644 --- a/Mathlib/Algebra/Order/Ring/Basic.lean +++ b/Mathlib/Algebra/Order/Ring/Basic.lean @@ -35,7 +35,7 @@ end MonoidHom section OrderedSemiring -variable [OrderedSemiring R] {a b x y : R} {n m : ℕ} +variable [OrderedSemiring R] {a b x y : R} {n : ℕ} theorem pow_add_pow_le (hx : 0 ≤ x) (hy : 0 ≤ y) (hn : n ≠ 0) : x ^ n + y ^ n ≤ (x + y) ^ n := by rcases Nat.exists_eq_add_one_of_ne_zero hn with ⟨k, rfl⟩ diff --git a/Mathlib/Algebra/Star/StarRingHom.lean b/Mathlib/Algebra/Star/StarRingHom.lean index 8d5fac3954d7d..539564a9df289 100644 --- a/Mathlib/Algebra/Star/StarRingHom.lean +++ b/Mathlib/Algebra/Star/StarRingHom.lean @@ -202,7 +202,7 @@ end Basic section Zero -- the `zero` requires extra type class assumptions because we need `star_zero` -variable {A B C D : Type*} +variable {A B C : Type*} variable [NonUnitalNonAssocSemiring A] [StarAddMonoid A] variable [NonUnitalNonAssocSemiring B] [StarAddMonoid B] @@ -425,7 +425,7 @@ variable {F G A B : Type*} variable [NonUnitalNonAssocSemiring A] [Star A] variable [NonUnitalNonAssocSemiring B] [Star B] variable [FunLike F A B] [NonUnitalRingHomClass F A B] [NonUnitalStarRingHomClass F A B] -variable [FunLike G B A] [NonUnitalRingHomClass G B A] [NonUnitalStarRingHomClass G B A] +variable [FunLike G B A] /-- If a (unital or non-unital) star ring morphism has an inverse, it is an isomorphism of star rings. -/ diff --git a/Mathlib/Analysis/Analytic/Within.lean b/Mathlib/Analysis/Analytic/Within.lean index b9f6a8bda1b70..67e0db70351e1 100644 --- a/Mathlib/Analysis/Analytic/Within.lean +++ b/Mathlib/Analysis/Analytic/Within.lean @@ -28,9 +28,8 @@ open Set Filter variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] -variable {E F G H : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] [NormedAddCommGroup F] - [NormedSpace 𝕜 F] [NormedAddCommGroup G] [NormedSpace 𝕜 G] [NormedAddCommGroup H] - [NormedSpace 𝕜 H] +variable {E F : Type*} + [NormedAddCommGroup E] [NormedSpace 𝕜 E] [NormedAddCommGroup F] [NormedSpace 𝕜 F] /-! ### Basic properties diff --git a/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean b/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean index f35fc2e27cedc..149c86f2c306e 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean @@ -111,13 +111,11 @@ attribute [local instance 1001] open Set Fin Filter Function -universe u uE uF uG uX +universe u uE uF variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] {E : Type uE} [NormedAddCommGroup E] - [NormedSpace 𝕜 E] {F : Type uF} [NormedAddCommGroup F] [NormedSpace 𝕜 F] {G : Type uG} - [NormedAddCommGroup G] [NormedSpace 𝕜 G] {X : Type uX} [NormedAddCommGroup X] [NormedSpace 𝕜 X] - {s s₁ t u : Set E} {f f₁ : E → F} {g : F → G} {x x₀ : E} {c : F} {m n : ℕ∞} - {p : E → FormalMultilinearSeries 𝕜 E F} + [NormedSpace 𝕜 E] {F : Type uF} [NormedAddCommGroup F] [NormedSpace 𝕜 F] + {s t u : Set E} {f f₁ : E → F} {x : E} {m n : ℕ∞} {p : E → FormalMultilinearSeries 𝕜 E F} /-! ### Functions with a Taylor series on a domain -/ diff --git a/Mathlib/Analysis/Calculus/FDeriv/Comp.lean b/Mathlib/Analysis/Calculus/FDeriv/Comp.lean index a185f886f2025..df6c6a6a2be03 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Comp.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Comp.lean @@ -30,12 +30,7 @@ variable {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] variable {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] variable {G : Type*} [NormedAddCommGroup G] [NormedSpace 𝕜 G] variable {G' : Type*} [NormedAddCommGroup G'] [NormedSpace 𝕜 G'] -variable {f f₀ f₁ g : E → F} -variable {f' f₀' f₁' g' : E →L[𝕜] F} -variable (e : E →L[𝕜] F) -variable {x : E} -variable {s t : Set E} -variable {L L₁ L₂ : Filter E} +variable {f g : E → F} {f' g' : E →L[𝕜] F} {x : E} {s : Set E} {L : Filter E} section Composition diff --git a/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean b/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean index e6343428b165f..f956fb2da11a4 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean @@ -34,12 +34,7 @@ variable {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] variable {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] variable {G : Type*} [NormedAddCommGroup G] [NormedSpace 𝕜 G] variable {G' : Type*} [NormedAddCommGroup G'] [NormedSpace 𝕜 G'] -variable {f f₀ f₁ g : E → F} -variable {f' f₀' f₁' g' : E →L[𝕜] F} -variable (e : E →L[𝕜] F) -variable {x : E} -variable {s t : Set E} -variable {L L₁ L₂ : Filter E} +variable {f : E → F} {f' : E →L[𝕜] F} {x : E} {s : Set E} namespace ContinuousLinearEquiv diff --git a/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean b/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean index 7de7ac1214369..2dfe679519c45 100644 --- a/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean +++ b/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean @@ -151,7 +151,7 @@ end namespace FormalMultilinearSeries variable [NontriviallyNormedField 𝕜] [NormedAddCommGroup E] [NormedSpace 𝕜 E] [NormedAddCommGroup F] - [NormedSpace 𝕜 F] [NormedAddCommGroup G] [NormedSpace 𝕜 G] + [NormedSpace 𝕜 F] variable (p : FormalMultilinearSeries 𝕜 E F) @@ -261,8 +261,8 @@ end Order section Coef -variable [NontriviallyNormedField 𝕜] [NormedAddCommGroup E] [NormedSpace 𝕜 E] {s : E} - {p : FormalMultilinearSeries 𝕜 𝕜 E} {f : 𝕜 → E} {n : ℕ} {z z₀ : 𝕜} {y : Fin n → 𝕜} +variable [NontriviallyNormedField 𝕜] [NormedAddCommGroup E] [NormedSpace 𝕜 E] + {p : FormalMultilinearSeries 𝕜 𝕜 E} {f : 𝕜 → E} {n : ℕ} {z : 𝕜} {y : Fin n → 𝕜} /-- The `n`th coefficient of `p` when seen as a power series. -/ def coeff (p : FormalMultilinearSeries 𝕜 𝕜 E) (n : ℕ) : E := diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean index b39e35ee29596..31e1077156f18 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean @@ -47,13 +47,11 @@ We use the following type variables in this file: universe u v v' wE wE₁ wE' wEi wG wG' -variable {𝕜 : Type u} {ι : Type v} {ι' : Type v'} {n : ℕ} {E : ι → Type wE} {E₁ : ι → Type wE₁} - {E' : ι' → Type wE'} {Ei : Fin n.succ → Type wEi} {G : Type wG} {G' : Type wG'} [Fintype ι] +variable {𝕜 : Type u} {ι : Type v} {ι' : Type v'} {n : ℕ} {E : ι → Type wE} + {Ei : Fin n.succ → Type wEi} {G : Type wG} {G' : Type wG'} [Fintype ι] [Fintype ι'] [NontriviallyNormedField 𝕜] [∀ i, NormedAddCommGroup (E i)] - [∀ i, NormedSpace 𝕜 (E i)] [∀ i, NormedAddCommGroup (E₁ i)] [∀ i, NormedSpace 𝕜 (E₁ i)] - [∀ i, NormedAddCommGroup (E' i)] [∀ i, NormedSpace 𝕜 (E' i)] [∀ i, NormedAddCommGroup (Ei i)] - [∀ i, NormedSpace 𝕜 (Ei i)] [NormedAddCommGroup G] [NormedSpace 𝕜 G] [NormedAddCommGroup G'] - [NormedSpace 𝕜 G'] + [∀ i, NormedSpace 𝕜 (E i)] [∀ i, NormedAddCommGroup (Ei i)] [∀ i, NormedSpace 𝕜 (Ei i)] + [NormedAddCommGroup G] [NormedSpace 𝕜 G] [NormedAddCommGroup G'] [NormedSpace 𝕜 G'] theorem ContinuousLinearMap.norm_map_tail_le diff --git a/Mathlib/Data/Finset/Insert.lean b/Mathlib/Data/Finset/Insert.lean index 94dd9b7abb2c2..1b93ae1f21cc9 100644 --- a/Mathlib/Data/Finset/Insert.lean +++ b/Mathlib/Data/Finset/Insert.lean @@ -43,7 +43,7 @@ open Multiset Subtype Function universe u -variable {α : Type*} {β : Type*} {γ : Type*} +variable {α : Type*} {β : Type*} namespace Finset @@ -345,7 +345,7 @@ end Cons section Insert -variable [DecidableEq α] {s t u v : Finset α} {a b : α} {f : α → β} +variable [DecidableEq α] {s t : Finset α} {a b : α} {f : α → β} /-- `insert a s` is the set `{a} ∪ s` containing `a` and the elements of `s`. -/ instance : Insert α (Finset α) := @@ -600,7 +600,7 @@ end Finset namespace Multiset -variable [DecidableEq α] {s t : Multiset α} +variable [DecidableEq α] @[simp] theorem toFinset_zero : toFinset (0 : Multiset α) = ∅ := @@ -618,8 +618,7 @@ end Multiset namespace List -variable [DecidableEq α] {l l' : List α} {a : α} {f : α → β} - {s : Finset α} {t : Set β} {t' : Finset β} +variable [DecidableEq α] {l : List α} {a : α} @[simp] theorem toFinset_nil : toFinset (@nil α) = ∅ := diff --git a/Mathlib/Data/Finset/Lattice/Basic.lean b/Mathlib/Data/Finset/Lattice/Basic.lean index c71d4b7a8faa4..aeb27e8cc2029 100644 --- a/Mathlib/Data/Finset/Lattice/Basic.lean +++ b/Mathlib/Data/Finset/Lattice/Basic.lean @@ -51,7 +51,7 @@ open Multiset Subtype Function universe u -variable {α : Type*} {β : Type*} {γ : Type*} +variable {α : Type*} namespace Finset @@ -62,7 +62,7 @@ attribute [local trans] Subset.trans Superset.trans section Lattice -variable [DecidableEq α] {s s₁ s₂ t t₁ t₂ u v : Finset α} {a b : α} +variable [DecidableEq α] {s s₁ s₂ t t₁ t₂ u v : Finset α} {a : α} /-- `s ∪ t` is the set such that `a ∈ s ∪ t` iff `a ∈ s` or `a ∈ t`. -/ instance : Union (Finset α) := diff --git a/Mathlib/Data/Matroid/Closure.lean b/Mathlib/Data/Matroid/Closure.lean index a8155c7f91667..4b7a3fddea385 100644 --- a/Mathlib/Data/Matroid/Closure.lean +++ b/Mathlib/Data/Matroid/Closure.lean @@ -283,7 +283,7 @@ lemma mem_closure_self (M : Matroid α) (e : α) (he : e ∈ M.E := by aesop_mat section Indep -variable {ι : Sort*} {I J B : Set α} {x y : α} +variable {ι : Sort*} {I J B : Set α} {x : α} lemma Indep.closure_eq_setOf_basis_insert (hI : M.Indep I) : M.closure I = {x | M.Basis I (insert x I)} := by diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index c3424288c710d..cebed32a1d12a 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -1449,7 +1449,7 @@ end DecidablePiExists section -variable [DecidableEq α] {s t u : Multiset α} {a b : α} +variable [DecidableEq α] {s t u : Multiset α} {a : α} /-- `s - t` is the multiset such that `count a (s - t) = count a s - count a t` for all `a` (note that it is truncated subtraction, so it is `0` if `count a t ≥ count a s`). -/ diff --git a/Mathlib/Data/Set/Functor.lean b/Mathlib/Data/Set/Functor.lean index 86ef698729203..1dbec4f1792df 100644 --- a/Mathlib/Data/Set/Functor.lean +++ b/Mathlib/Data/Set/Functor.lean @@ -20,7 +20,7 @@ open Function Set.Notation namespace Set -variable {α β : Type u} {s : Set α} {f : α → Set β} {g : Set (α → β)} +variable {α β : Type u} {s : Set α} {f : α → Set β} /-- The `Set` functor is a monad. diff --git a/Mathlib/Data/Set/Sigma.lean b/Mathlib/Data/Set/Sigma.lean index 7201429b16acc..b8553b4a83683 100644 --- a/Mathlib/Data/Set/Sigma.lean +++ b/Mathlib/Data/Set/Sigma.lean @@ -14,7 +14,7 @@ This file defines `Set.sigma`, the indexed sum of sets. namespace Set -variable {ι ι' : Type*} {α β : ι → Type*} {s s₁ s₂ : Set ι} {t t₁ t₂ : ∀ i, Set (α i)} +variable {ι ι' : Type*} {α : ι → Type*} {s s₁ s₂ : Set ι} {t t₁ t₂ : ∀ i, Set (α i)} {u : Set (Σ i, α i)} {x : Σ i, α i} {i j : ι} {a : α i} @[simp] diff --git a/Mathlib/Deprecated/Submonoid.lean b/Mathlib/Deprecated/Submonoid.lean index f85c670a230b8..639dc73b90592 100644 --- a/Mathlib/Deprecated/Submonoid.lean +++ b/Mathlib/Deprecated/Submonoid.lean @@ -29,8 +29,7 @@ Submonoid, Submonoids, IsSubmonoid -/ -variable {M : Type*} [Monoid M] {s : Set M} -variable {A : Type*} [AddMonoid A] {t : Set A} +variable {M : Type*} [Monoid M] {s : Set M} {A : Type*} [AddMonoid A] /-- `s` is an additive submonoid: a set containing 0 and closed under addition. Note that this structure is deprecated, and the bundled variant `AddSubmonoid A` should be diff --git a/Mathlib/ModelTheory/DirectLimit.lean b/Mathlib/ModelTheory/DirectLimit.lean index 7744cba16c693..d262cae6955b5 100644 --- a/Mathlib/ModelTheory/DirectLimit.lean +++ b/Mathlib/ModelTheory/DirectLimit.lean @@ -441,7 +441,7 @@ end DirectLimit section Substructure variable [Nonempty ι] [IsDirected ι (· ≤ ·)] -variable {M N : Type*} [L.Structure M] [L.Structure N] (S : ι →o L.Substructure M) +variable {M : Type*} [L.Structure M] (S : ι →o L.Substructure M) instance : DirectedSystem (fun i ↦ S i) (fun _ _ h ↦ Substructure.inclusion (S.monotone h)) where map_self _ _ := rfl diff --git a/Mathlib/NumberTheory/ModularForms/SlashActions.lean b/Mathlib/NumberTheory/ModularForms/SlashActions.lean index 3103cf011795b..4b39fc4a7a847 100644 --- a/Mathlib/NumberTheory/ModularForms/SlashActions.lean +++ b/Mathlib/NumberTheory/ModularForms/SlashActions.lean @@ -76,7 +76,7 @@ noncomputable section def slash (k : ℤ) (γ : GL(2, ℝ)⁺) (f : ℍ → ℂ) (x : ℍ) : ℂ := f (γ • x) * (↑(↑ₘ[ℝ] γ).det : ℂ) ^ (k - 1) * UpperHalfPlane.denom γ x ^ (-k) -variable {Γ : Subgroup SL(2, ℤ)} {k : ℤ} (f : ℍ → ℂ) +variable {k : ℤ} (f : ℍ → ℂ) section diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean index 01eff7db3e0be..bf1aa9e9ffaec 100644 --- a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean +++ b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean @@ -661,7 +661,7 @@ instance Pi.conditionallyCompleteLattice {ι : Type*} {α : ι → Type*} section ConditionallyCompleteLinearOrder -variable [ConditionallyCompleteLinearOrder α] {s t : Set α} {a b : α} +variable [ConditionallyCompleteLinearOrder α] {s : Set α} {a b : α} /-- When `b < sSup s`, there is an element `a` in `s` with `b < a`, if `s` is nonempty and the order is a linear order. -/ @@ -790,7 +790,7 @@ section ConditionallyCompleteLinearOrderBot theorem csInf_univ [ConditionallyCompleteLinearOrder α] [OrderBot α] : sInf (univ : Set α) = ⊥ := isLeast_univ.csInf_eq -variable [ConditionallyCompleteLinearOrderBot α] {s : Set α} {f : ι → α} {a : α} +variable [ConditionallyCompleteLinearOrderBot α] {s : Set α} {a : α} @[simp] theorem csSup_empty : (sSup ∅ : α) = ⊥ := @@ -1015,7 +1015,7 @@ a binary function whose partial evaluations are lower/upper adjoints of Galois c section variable [ConditionallyCompleteLattice α] [ConditionallyCompleteLattice β] - [ConditionallyCompleteLattice γ] {f : α → β → γ} {s : Set α} {t : Set β} + [ConditionallyCompleteLattice γ] {s : Set α} {t : Set β} variable {l u : α → β → γ} {l₁ u₁ : β → γ → α} {l₂ u₂ : α → γ → β} diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Indexed.lean b/Mathlib/Order/ConditionallyCompleteLattice/Indexed.lean index 667d696dff1b2..1d69aba7a1e11 100644 --- a/Mathlib/Order/ConditionallyCompleteLattice/Indexed.lean +++ b/Mathlib/Order/ConditionallyCompleteLattice/Indexed.lean @@ -61,7 +61,7 @@ end section ConditionallyCompleteLattice -variable [ConditionallyCompleteLattice α] {s t : Set α} {a b : α} +variable [ConditionallyCompleteLattice α] {a b : α} theorem isLUB_ciSup [Nonempty ι] {f : ι → α} (H : BddAbove (range f)) : IsLUB (range f) (⨆ i, f i) := @@ -354,7 +354,7 @@ end ConditionallyCompleteLattice section ConditionallyCompleteLinearOrder -variable [ConditionallyCompleteLinearOrder α] {s t : Set α} {a b : α} +variable [ConditionallyCompleteLinearOrder α] {a b : α} /-- Indexed version of `exists_lt_of_lt_csSup`. When `b < iSup f`, there is an element `i` such that `b < f i`. @@ -444,7 +444,7 @@ In this case we have `Sup ∅ = ⊥`, so we can drop some `Nonempty`/`Set.Nonemp section ConditionallyCompleteLinearOrderBot -variable [ConditionallyCompleteLinearOrderBot α] {s : Set α} {f : ι → α} {a : α} +variable [ConditionallyCompleteLinearOrderBot α] {f : ι → α} {a : α} @[simp] theorem ciSup_of_empty [IsEmpty ι] (f : ι → α) : ⨆ i, f i = ⊥ := by diff --git a/Mathlib/Order/Interval/Finset/Box.lean b/Mathlib/Order/Interval/Finset/Box.lean index 902ac28210bbd..1d5823025b14e 100644 --- a/Mathlib/Order/Interval/Finset/Box.lean +++ b/Mathlib/Order/Interval/Finset/Box.lean @@ -77,7 +77,7 @@ end Prod /-! ### `ℤ × ℤ` -/ namespace Int -variable {n : ℕ} {x : ℤ × ℤ} +variable {x : ℤ × ℤ} attribute [norm_cast] toNat_ofNat diff --git a/Mathlib/Order/Interval/Set/UnorderedInterval.lean b/Mathlib/Order/Interval/Set/UnorderedInterval.lean index 68a76e72afd03..b95aff755afaa 100644 --- a/Mathlib/Order/Interval/Set/UnorderedInterval.lean +++ b/Mathlib/Order/Interval/Set/UnorderedInterval.lean @@ -339,7 +339,7 @@ lemma mem_uIoo_of_lt (ha : a < x) (hb : x < b) : x ∈ uIoo a b := Ioo_subset_uI lemma mem_uIoo_of_gt (hb : b < x) (ha : x < a) : x ∈ uIoo a b := Ioo_subset_uIoo' ⟨hb, ha⟩ -variable [LinearOrder β] {f : α → β} {s : Set α} {a a₁ a₂ b b₁ b₂ c d x : α} +variable {a b : α} theorem Ioo_min_max : Ioo (min a b) (max a b) = uIoo a b := rfl diff --git a/Mathlib/RingTheory/Derivation/Lie.lean b/Mathlib/RingTheory/Derivation/Lie.lean index ba28f09e9c364..cbd9e3c53761e 100644 --- a/Mathlib/RingTheory/Derivation/Lie.lean +++ b/Mathlib/RingTheory/Derivation/Lie.lean @@ -18,7 +18,7 @@ namespace Derivation variable {R : Type*} [CommRing R] variable {A : Type*} [CommRing A] [Algebra R A] -variable (D : Derivation R A A) {D1 D2 : Derivation R A A} (a : A) +variable {D1 D2 : Derivation R A A} (a : A) section LieStructures diff --git a/Mathlib/Topology/Separation/GDelta.lean b/Mathlib/Topology/Separation/GDelta.lean index 9ed1577b0cb03..739f641064df5 100644 --- a/Mathlib/Topology/Separation/GDelta.lean +++ b/Mathlib/Topology/Separation/GDelta.lean @@ -26,9 +26,9 @@ occasionally the literature swaps definitions for e.g. T₃ and regular. open Function Set Filter Topology TopologicalSpace -universe u v +universe u -variable {X : Type*} {Y : Type*} [TopologicalSpace X] +variable {X : Type*} [TopologicalSpace X] section Separation From 118a685fff56f220efc657cd0d0dd7a13dfdf3d8 Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Sat, 16 Nov 2024 00:52:44 +0000 Subject: [PATCH 005/829] feat(FieldTheory/SeparableDegree): `Field.Emb` is infinite for transcendental extension (#18770) - `Field.infinite_emb_of_transcendental`, `Field.finSepDegree_eq_zero_of_transcendental`: `Field.Emb` is infinite and `Field.finSepDegree` is zero for transcendental extensions - remove algebraic assumptions in `isPurelyInseparable_(of|iff)_finSepDegree_eq_one` - prove `IsPurelyInseparable.of_injective_comp_algebraMap` - change `Field.Emb` from `def` to `abbrev` --- Mathlib/FieldTheory/PurelyInseparable.lean | 58 +++++++++++----------- Mathlib/FieldTheory/SeparableDegree.lean | 33 +++++++++++- 2 files changed, 61 insertions(+), 30 deletions(-) diff --git a/Mathlib/FieldTheory/PurelyInseparable.lean b/Mathlib/FieldTheory/PurelyInseparable.lean index af9e58badb134..f077e7c7ee4e4 100644 --- a/Mathlib/FieldTheory/PurelyInseparable.lean +++ b/Mathlib/FieldTheory/PurelyInseparable.lean @@ -53,11 +53,9 @@ of fields. characteristic `q` is purely inseparable if and only if for every element `x` of `E`, the minimal polynomial of `x` over `F` is of form `(X - x) ^ (q ^ n)` for some natural number `n`. -- `isPurelyInseparable_iff_finSepDegree_eq_one`: an algebraic extension is purely inseparable +- `isPurelyInseparable_iff_finSepDegree_eq_one`: an extension is purely inseparable if and only if it has finite separable degree (`Field.finSepDegree`) one. - **TODO:** remove the algebraic assumption. - - `IsPurelyInseparable.normal`: a purely inseparable extension is normal. - `separableClosure.isPurelyInseparable`: if `E / F` is algebraic, then `E` is purely inseparable @@ -80,6 +78,11 @@ of fields. reduced ring `L`, the map `(E →+* L) → (F →+* L)` induced by `algebraMap F E` is injective. In particular, a purely inseparable field extension is an epimorphism in the category of fields. +- `IsPurelyInseparable.of_injective_comp_algebraMap`: if `L` is an algebraically closed field + containing `E`, such that the map `(E →+* L) → (F →+* L)` induced by `algebraMap F E` is + injective, then `E / F` is purely inseparable. As a corollary, epimorphisms in the category of + fields must be purely inseparable extensions. + - `IntermediateField.isPurelyInseparable_adjoin_iff_pow_mem`: if `F` is of exponential characteristic `q`, then `F(S) / F` is a purely inseparable extension if and only if for any `x ∈ S`, `x ^ (q ^ n)` is contained in `F` for some `n : ℕ`. @@ -115,12 +118,6 @@ separable degree, degree, separable closure, purely inseparable ## TODO -- `IsPurelyInseparable.of_injective_comp_algebraMap`: if `L` is an algebraically closed field - containing `E`, such that the map `(E →+* L) → (F →+* L)` induced by `algebraMap F E` is - injective, then `E / F` is purely inseparable. As a corollary, epimorphisms in the category of - fields must be purely inseparable extensions. Need to use the fact that `Emb F E` is infinite - (or just not a singleton) when `E / F` is (purely) transcendental. - - Restate some intermediate result in terms of linearly disjointness. - Prove that the inseparable degrees satisfy the tower law: $[E:F]_i [K:E]_i = [K:F]_i$. @@ -511,19 +508,21 @@ theorem IsPurelyInseparable.minpoly_eq_X_sub_C_pow (q : ℕ) [ExpChar F q] [IsPu variable (E) --- TODO: remove `halg` assumption variable {F E} in -/-- If an algebraic extension has finite separable degree one, then it is purely inseparable. -/ -theorem isPurelyInseparable_of_finSepDegree_eq_one [Algebra.IsAlgebraic F E] +/-- If an extension has finite separable degree one, then it is purely inseparable. -/ +theorem isPurelyInseparable_of_finSepDegree_eq_one (hdeg : finSepDegree F E = 1) : IsPurelyInseparable F E := by - rw [isPurelyInseparable_iff] - refine fun x ↦ ⟨Algebra.IsIntegral.isIntegral x, fun hsep ↦ ?_⟩ - have : Algebra.IsAlgebraic F⟮x⟯ E := Algebra.IsAlgebraic.tower_top (K := F) F⟮x⟯ - have := finSepDegree_mul_finSepDegree_of_isAlgebraic F F⟮x⟯ E - rw [hdeg, mul_eq_one, (finSepDegree_adjoin_simple_eq_finrank_iff F E x - (Algebra.IsAlgebraic.isAlgebraic x)).2 hsep, - IntermediateField.finrank_eq_one_iff] at this - simpa only [this.1] using mem_adjoin_simple_self F x + by_cases H : Algebra.IsAlgebraic F E + · rw [isPurelyInseparable_iff] + refine fun x ↦ ⟨Algebra.IsIntegral.isIntegral x, fun hsep ↦ ?_⟩ + have : Algebra.IsAlgebraic F⟮x⟯ E := Algebra.IsAlgebraic.tower_top (K := F) F⟮x⟯ + have := finSepDegree_mul_finSepDegree_of_isAlgebraic F F⟮x⟯ E + rw [hdeg, mul_eq_one, (finSepDegree_adjoin_simple_eq_finrank_iff F E x + (Algebra.IsAlgebraic.isAlgebraic x)).2 hsep, + IntermediateField.finrank_eq_one_iff] at this + simpa only [this.1] using mem_adjoin_simple_self F x + · rw [← Algebra.transcendental_iff_not_isAlgebraic] at H + simp [finSepDegree_eq_zero_of_transcendental F E] at hdeg namespace IsPurelyInseparable @@ -600,10 +599,8 @@ theorem IsPurelyInseparable.insepDegree_eq [IsPurelyInseparable F E] : theorem IsPurelyInseparable.finInsepDegree_eq [IsPurelyInseparable F E] : finInsepDegree F E = finrank F E := congr(Cardinal.toNat $(insepDegree_eq F E)) --- TODO: remove `halg` assumption -/-- An algebraic extension is purely inseparable if and only if it has finite separable -degree one. -/ -theorem isPurelyInseparable_iff_finSepDegree_eq_one [Algebra.IsAlgebraic F E] : +/-- An extension is purely inseparable if and only if it has finite separable degree one. -/ +theorem isPurelyInseparable_iff_finSepDegree_eq_one : IsPurelyInseparable F E ↔ finSepDegree F E = 1 := ⟨fun _ ↦ IsPurelyInseparable.finSepDegree_eq_one F E, fun h ↦ isPurelyInseparable_of_finSepDegree_eq_one h⟩ @@ -683,15 +680,18 @@ theorem eq_separableClosure_iff [Algebra.IsAlgebraic F E] (L : IntermediateField ⟨by rintro rfl; exact ⟨isSeparable F E, isPurelyInseparable F E⟩, fun ⟨_, _⟩ ↦ eq_separableClosure F E L⟩ --- TODO: prove it -set_option linter.unusedVariables false in /-- If `L` is an algebraically closed field containing `E`, such that the map `(E →+* L) → (F →+* L)` induced by `algebraMap F E` is injective, then `E / F` is purely inseparable. As a corollary, epimorphisms in the category of fields must be purely inseparable extensions. -/ -proof_wanted IsPurelyInseparable.of_injective_comp_algebraMap (L : Type w) [Field L] [IsAlgClosed L] - (hn : Nonempty (E →+* L)) (h : Function.Injective fun f : E →+* L ↦ f.comp (algebraMap F E)) : - IsPurelyInseparable F E +theorem IsPurelyInseparable.of_injective_comp_algebraMap (L : Type w) [Field L] [IsAlgClosed L] + [Nonempty (E →+* L)] (h : Function.Injective fun f : E →+* L ↦ f.comp (algebraMap F E)) : + IsPurelyInseparable F E := by + rw [isPurelyInseparable_iff_finSepDegree_eq_one, finSepDegree, Nat.card_eq_one_iff_unique] + letI := (Classical.arbitrary (E →+* L)).toAlgebra + let j : AlgebraicClosure E →ₐ[E] L := IsAlgClosed.lift + exact ⟨⟨fun f g ↦ DFunLike.ext' <| j.injective.comp_left (congr_arg (⇑) <| + @h (j.toRingHom.comp f) (j.toRingHom.comp g) (by ext; simp))⟩, inferInstance⟩ end IsPurelyInseparable diff --git a/Mathlib/FieldTheory/SeparableDegree.lean b/Mathlib/FieldTheory/SeparableDegree.lean index 23f7b5124d475..8025c4eff3c76 100644 --- a/Mathlib/FieldTheory/SeparableDegree.lean +++ b/Mathlib/FieldTheory/SeparableDegree.lean @@ -7,6 +7,7 @@ import Mathlib.FieldTheory.SplittingField.Construction import Mathlib.FieldTheory.IsAlgClosed.AlgebraicClosure import Mathlib.FieldTheory.Separable import Mathlib.FieldTheory.NormalClosure +import Mathlib.RingTheory.AlgebraicIndependent import Mathlib.RingTheory.Polynomial.SeparableDegree import Mathlib.RingTheory.Polynomial.UniqueFactorization @@ -70,6 +71,8 @@ This file contains basics about the separable degree of a field extension. In particular, the separable degrees satisfy the tower law: $[E:F]_s [K:E]_s = [K:F]_s$ (see also `Module.finrank_mul_finrank`). +- `Field.infinite_emb_of_transcendental`: `Field.Emb` is infinite for transcendental extensions. + - `Polynomial.natSepDegree_le_natDegree`: the separable degree of a polynomial is smaller than its degree. @@ -132,7 +135,7 @@ namespace Field /-- `Field.Emb F E` is the type of `F`-algebra homomorphisms from `E` to the algebraic closure of `E`. -/ -def Emb := E →ₐ[F] AlgebraicClosure E +abbrev Emb := E →ₐ[F] AlgebraicClosure E /-- If `E / F` is an algebraic extension, then the (finite) separable degree of `E / F` is the number of `F`-algebra homomorphisms from `E` to the algebraic closure of `E`, @@ -246,6 +249,34 @@ def embProdEmbOfIsAlgebraic [Algebra E K] [IsScalarTower F E K] [Algebra.IsAlgeb (IsAlgClosure.equivOfAlgebraic E K (AlgebraicClosure K) (AlgebraicClosure E)).restrictScalars F).symm +/-- If the field extension `E / F` is transcendental, then `Field.Emb F E` is infinite. -/ +instance infinite_emb_of_transcendental [H : Algebra.Transcendental F E] : Infinite (Emb F E) := by + obtain ⟨ι, x, hx⟩ := exists_isTranscendenceBasis' _ (algebraMap F E).injective + have := hx.isAlgebraic_field + rw [← (embProdEmbOfIsAlgebraic F (adjoin F (Set.range x)) E).infinite_iff] + refine @Prod.infinite_of_left _ _ ?_ _ + rw [← (embEquivOfEquiv _ _ _ hx.1.aevalEquivField).infinite_iff] + obtain ⟨i⟩ := hx.nonempty_iff_transcendental.2 H + let K := FractionRing (MvPolynomial ι F) + let i1 := IsScalarTower.toAlgHom F (MvPolynomial ι F) (AlgebraicClosure K) + have hi1 : Function.Injective i1 := by + rw [IsScalarTower.coe_toAlgHom', IsScalarTower.algebraMap_eq _ K] + exact (algebraMap K (AlgebraicClosure K)).injective.comp (IsFractionRing.injective _ _) + let f (n : ℕ) : Emb F K := IsFractionRing.liftAlgHom + (g := i1.comp <| MvPolynomial.aeval fun i : ι ↦ MvPolynomial.X i ^ (n + 1)) <| hi1.comp <| by + simpa [algebraicIndependent_iff_injective_aeval] using + MvPolynomial.algebraicIndependent_polynomial_aeval_X _ + fun i : ι ↦ (Polynomial.transcendental_X F).pow n.succ_pos + refine Infinite.of_injective f fun m n h ↦ ?_ + replace h : (MvPolynomial.X i) ^ (m + 1) = (MvPolynomial.X i) ^ (n + 1) := hi1 <| by + simpa [f, -map_pow] using congr($h (algebraMap _ K (MvPolynomial.X (R := F) i))) + simpa using congr(MvPolynomial.totalDegree $h) + +/-- If the field extension `E / F` is transcendental, then `Field.finSepDegree F E = 0`, which +actually means that `Field.Emb F E` is infinite (see `Field.infinite_emb_of_transcendental`). -/ +theorem finSepDegree_eq_zero_of_transcendental [Algebra.Transcendental F E] : + finSepDegree F E = 0 := Nat.card_eq_zero_of_infinite + /-- 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 `Module.finrank_mul_finrank`. -/ From 46215793fb6b682b6eeae1960e3266f0aadd5794 Mon Sep 17 00:00:00 2001 From: Alex Kontorovich <58564076+AlexKontorovich@users.noreply.github.com> Date: Sat, 16 Nov 2024 02:46:43 +0000 Subject: [PATCH 006/829] feat: `continuousAt_iff_isLittleO` (#19051) This is another PR breaking #9598 into smaller pieces. We prove the theorem `continuousAt_iff_isLittleO`, that a function `f` is continuous at a point `x` iff `f y - f x` is little-o of `1`. --- Mathlib/Analysis/Asymptotics/Asymptotics.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Mathlib/Analysis/Asymptotics/Asymptotics.lean b/Mathlib/Analysis/Asymptotics/Asymptotics.lean index 3c099c0ae33cb..87c9e471f1eac 100644 --- a/Mathlib/Analysis/Asymptotics/Asymptotics.lean +++ b/Mathlib/Analysis/Asymptotics/Asymptotics.lean @@ -1254,6 +1254,11 @@ theorem IsLittleO.trans_tendsto (hfg : f'' =o[l] g'') (hg : Tendsto g'' l (𝓝 lemma isLittleO_id_one [One F''] [NeZero (1 : F'')] : (fun x : E'' => x) =o[𝓝 0] (1 : E'' → F'') := isLittleO_id_const one_ne_zero +theorem continuousAt_iff_isLittleO {α : Type*} {E : Type*} [NormedRing E] [NormOneClass E] + [TopologicalSpace α] {f : α → E} {x : α} : + (ContinuousAt f x) ↔ (fun (y : α) ↦ f y - f x) =o[𝓝 x] (fun (_ : α) ↦ (1 : E)) := by + simp [ContinuousAt, ← tendsto_sub_nhds_zero_iff] + /-! ### Multiplication by a constant -/ From aa5b6a4acd12e898badba2f2cf7aaaff253feb20 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Sat, 16 Nov 2024 06:58:45 +0000 Subject: [PATCH 007/829] feat(AlgebraicGeometry): define affine `n`-space (#18837) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/AlgebraicGeometry/AffineScheme.lean | 4 + Mathlib/AlgebraicGeometry/AffineSpace.lean | 386 ++++++++++++++++++++ 3 files changed, 391 insertions(+) create mode 100644 Mathlib/AlgebraicGeometry/AffineSpace.lean diff --git a/Mathlib.lean b/Mathlib.lean index 7e8dc34a4921f..64355d7319159 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -907,6 +907,7 @@ import Mathlib.Algebra.Tropical.BigOperators import Mathlib.Algebra.Tropical.Lattice import Mathlib.Algebra.Vertex.HVertexOperator import Mathlib.AlgebraicGeometry.AffineScheme +import Mathlib.AlgebraicGeometry.AffineSpace import Mathlib.AlgebraicGeometry.Cover.MorphismProperty import Mathlib.AlgebraicGeometry.Cover.Open import Mathlib.AlgebraicGeometry.EllipticCurve.Affine diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index bce99104712f3..956a2db40bc22 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -134,6 +134,10 @@ theorem Scheme.isoSpec_Spec (R : CommRingCat.{u}) : (Spec R).isoSpec.inv = Spec.map (Scheme.ΓSpecIso R).inv := congr($(isoSpec_Spec R).inv) +lemma ext_of_isAffine {X Y : Scheme} [IsAffine Y] {f g : X ⟶ Y} (e : f.app ⊤ = g.app ⊤) : + f = g := by + rw [← cancel_mono Y.toSpecΓ, Scheme.toSpecΓ_naturality, Scheme.toSpecΓ_naturality, e] + namespace AffineScheme /-- The `Spec` functor into the category of affine schemes. -/ diff --git a/Mathlib/AlgebraicGeometry/AffineSpace.lean b/Mathlib/AlgebraicGeometry/AffineSpace.lean new file mode 100644 index 0000000000000..6374a6564ad9f --- /dev/null +++ b/Mathlib/AlgebraicGeometry/AffineSpace.lean @@ -0,0 +1,386 @@ +/- +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.Algebra.MvPolynomial.Monad +import Mathlib.AlgebraicGeometry.Limits + +/-! +# Affine space + +## Main definitions + +- `AlgebraicGeometry.AffineSpace`: `𝔸(n; S)` is the affine `n`-space over `S`. +- `AlgebraicGeometry.AffineSpace.coord`: The standard coordinate functions on the affine space. +- `AlgebraicGeometry.AffineSpace.homOfVector`: + The morphism `X ⟶ 𝔸(n; S)` given by a `X ⟶ S` and a choice of `n`-coordinate functions. +- `AlgebraicGeometry.AffineSpace.homOverEquiv`: + `S`-morphisms into `Spec 𝔸(n; S)` are equivalent to the choice of `n` global sections. +- `AlgebraicGeometry.AffineSpace.SpecIso`: `𝔸(n; Spec R) ≅ Spec R[n]` + +-/ + +open CategoryTheory Limits MvPolynomial + +noncomputable section + +namespace AlgebraicGeometry + +universe v u + +variable (n : Type v) (S : Scheme.{max u v}) + +local notation3 "ℤ[" n "]" => CommRingCat.of (MvPolynomial n (ULift ℤ)) +local notation3 "ℤ[" n "].{" u "}" => CommRingCat.of (MvPolynomial n (ULift.{u} ℤ)) + +/-- `𝔸(n; S)` is the affine `n`-space over `S`. +Note that `n` is an arbitrary index type (e.g. `Fin m`). -/ +def AffineSpace (n : Type v) (S : Scheme.{max u v}) : Scheme.{max u v} := + pullback (terminal.from S) (terminal.from (Spec ℤ[n])) + +namespace AffineSpace + +/-- `𝔸(n; S)` is the affine `n`-space over `S`. -/ +scoped [AlgebraicGeometry] notation "𝔸("n"; "S")" => AffineSpace n S + +variable {n} in +lemma of_mvPolynomial_int_ext {R} {f g : ℤ[n] ⟶ R} (h : ∀ i, f (.X i) = g (.X i)) : f = g := by + suffices f.comp (MvPolynomial.mapEquiv _ ULift.ringEquiv.symm).toRingHom = + g.comp (MvPolynomial.mapEquiv _ ULift.ringEquiv.symm).toRingHom by + ext x + obtain ⟨x, rfl⟩ := (MvPolynomial.mapEquiv _ ULift.ringEquiv.symm).surjective x + exact DFunLike.congr_fun this x + ext1 + · exact RingHom.ext_int _ _ + · simpa using h _ + + +@[simps (config := .lemmasOnly)] +instance over : 𝔸(n; S).CanonicallyOver S where + hom := pullback.fst _ _ + +/-- The map from the affine `n`-space over `S` to the integral model `Spec ℤ[n]`. -/ +def toSpecMvPoly : 𝔸(n; S) ⟶ Spec ℤ[n] := pullback.snd _ _ + +variable {X : Scheme.{max u v}} + +/-- +Morphisms into `Spec ℤ[n]` are equivalent the choice of `n` global sections. +Use `homOverEquiv` instead. +-/ +@[simps] +def toSpecMvPolyIntEquiv : (X ⟶ Spec ℤ[n]) ≃ (n → Γ(X, ⊤)) where + toFun f i := f.app ⊤ ((Scheme.ΓSpecIso ℤ[n]).inv (.X i)) + invFun v := X.toSpecΓ ≫ Spec.map + (MvPolynomial.eval₂Hom ((algebraMap ℤ _).comp ULift.ringEquiv.toRingHom) v) + left_inv f := by + apply (ΓSpec.adjunction.homEquiv _ _).symm.injective + apply Quiver.Hom.unop_inj + rw [Adjunction.homEquiv_symm_apply, Adjunction.homEquiv_symm_apply] + simp only [Functor.rightOp_obj, Scheme.Γ_obj, Scheme.Spec_obj, algebraMap_int_eq, + RingEquiv.toRingHom_eq_coe, TopologicalSpace.Opens.map_top, Functor.rightOp_map, op_comp, + Scheme.Γ_map, unop_comp, Quiver.Hom.unop_op, Scheme.comp_app, Scheme.toSpecΓ_app_top, + Scheme.ΓSpecIso_naturality, ΓSpec.adjunction_counit_app, Category.assoc, + Iso.cancel_iso_inv_left, ← Iso.eq_inv_comp] + apply of_mvPolynomial_int_ext + intro i + rw [coe_eval₂Hom, eval₂_X] + rfl + right_inv v := by + ext i + simp only [algebraMap_int_eq, RingEquiv.toRingHom_eq_coe, Scheme.comp_coeBase, + TopologicalSpace.Opens.map_comp_obj, TopologicalSpace.Opens.map_top, Scheme.comp_app, + Scheme.toSpecΓ_app_top, Scheme.ΓSpecIso_naturality, CommRingCat.comp_apply, + CommRingCat.coe_of] + erw [Iso.hom_inv_id_apply] + rw [coe_eval₂Hom, eval₂_X] + +lemma toSpecMvPolyIntEquiv_comp {X Y : Scheme} (f : X ⟶ Y) (g : Y ⟶ Spec ℤ[n]) (i) : + toSpecMvPolyIntEquiv n (f ≫ g) i = f.app ⊤ (toSpecMvPolyIntEquiv n g i) := rfl + +variable {n} in +/-- The standard coordinates of `𝔸(n; S)`. -/ +def coord (i : n) : Γ(𝔸(n; S), ⊤) := toSpecMvPolyIntEquiv _ (toSpecMvPoly n S) i + +section homOfVector + +variable {n S} + +/-- The morphism `X ⟶ 𝔸(n; S)` given by a `X ⟶ S` and a choice of `n`-coordinate functions. -/ +def homOfVector (f : X ⟶ S) (v : n → Γ(X, ⊤)) : X ⟶ 𝔸(n; S) := + pullback.lift f ((toSpecMvPolyIntEquiv n).symm v) (by simp) + +variable (f : X ⟶ S) (v : n → Γ(X, ⊤)) + +@[reassoc (attr := simp)] +lemma homOfVector_over : homOfVector f v ≫ 𝔸(n; S) ↘ S = f := + pullback.lift_fst _ _ _ + +@[reassoc] +lemma homOfVector_toSpecMvPoly : + homOfVector f v ≫ toSpecMvPoly n S = (toSpecMvPolyIntEquiv n).symm v := + pullback.lift_snd _ _ _ + +@[local simp] +lemma homOfVector_app_top_coord (i) : + (homOfVector f v).app ⊤ (coord S i) = v i := by + rw [coord, ← toSpecMvPolyIntEquiv_comp, homOfVector_toSpecMvPoly, + Equiv.apply_symm_apply] + +@[ext 1100] +lemma hom_ext {f g : X ⟶ 𝔸(n; S)} + (h₁ : f ≫ 𝔸(n; S) ↘ S = g ≫ 𝔸(n; S) ↘ S) + (h₂ : ∀ i, f.app ⊤ (coord S i) = g.app ⊤ (coord S i)) : f = g := by + apply pullback.hom_ext h₁ + show f ≫ toSpecMvPoly _ _ = g ≫ toSpecMvPoly _ _ + apply (toSpecMvPolyIntEquiv n).injective + ext i + rw [toSpecMvPolyIntEquiv_comp, toSpecMvPolyIntEquiv_comp] + exact h₂ i + +@[reassoc] +lemma comp_homOfVector {X Y : Scheme} (v : n → Γ(Y, ⊤)) (f : X ⟶ Y) (g : Y ⟶ S) : + f ≫ homOfVector g v = homOfVector (f ≫ g) (f.app ⊤ ∘ v) := by + ext1 <;> simp [-TopologicalSpace.Opens.map_top]; rfl + +end homOfVector + +variable [X.Over S] + +variable {n} + +instance (v : n → Γ(X, ⊤)) : (homOfVector (X ↘ S) v).IsOver S where + +/-- `S`-morphisms into `Spec 𝔸(n; S)` are equivalent to the choice of `n` global sections. -/ +@[simps] +def homOverEquiv : { f : X ⟶ 𝔸(n; S) // f.IsOver S } ≃ (n → Γ(X, ⊤)) where + toFun f i := f.1.app ⊤ (coord S i) + invFun v := ⟨homOfVector (X ↘ S) v, inferInstance⟩ + left_inv f := by + ext : 2 + · simp [f.2.1] + · rw [homOfVector_app_top_coord] + right_inv v := by ext i; simp [-TopologicalSpace.Opens.map_top, homOfVector_app_top_coord] + +variable (n) in +/-- +The affine space over an affine base is isomorphic to the spectrum of the polynomial ring. +Also see `AffineSpace.SpecIso`. +-/ +@[simps (config := .lemmasOnly) hom inv] +def isoOfIsAffine [IsAffine S] : + 𝔸(n; S) ≅ Spec (.of (MvPolynomial n Γ(S, ⊤))) where + hom := 𝔸(n; S).toSpecΓ ≫ Spec.map (eval₂Hom ((𝔸(n; S) ↘ S).app ⊤) (coord S)) + inv := homOfVector (Spec.map C ≫ S.isoSpec.inv) + ((Scheme.ΓSpecIso (.of (MvPolynomial n Γ(S, ⊤)))).inv ∘ MvPolynomial.X) + hom_inv_id := by + ext1 + · simp only [Category.assoc, homOfVector_over, Category.id_comp] + rw [← Spec.map_comp_assoc, CommRingCat.comp_eq_ring_hom_comp, eval₂Hom_comp_C, + ← Scheme.toSpecΓ_naturality_assoc] + simp [Scheme.isoSpec] + · simp only [Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, Scheme.comp_app, + CommRingCat.coe_comp_of, RingHom.coe_comp, Function.comp_apply, Scheme.id.base, + Scheme.id_app] + -- note: `rw [homOfVector_app_top_coord]` works but takes 3 more seconds + conv => enter [1, 2, 2]; rw [homOfVector_app_top_coord] + simp only [TopologicalSpace.Opens.map_top, Scheme.toSpecΓ_app_top, Function.comp_apply, + CommRingCat.id_apply] + erw [elementwise_of% Scheme.ΓSpecIso_naturality, Iso.hom_inv_id_apply] + exact eval₂_X _ _ _ + inv_hom_id := by + apply ext_of_isAffine + simp only [Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, + TopologicalSpace.Opens.map_top, Scheme.comp_app, Scheme.toSpecΓ_app_top, + Scheme.ΓSpecIso_naturality, Category.assoc, Scheme.id_app, ← Iso.eq_inv_comp, + Category.comp_id] + apply ringHom_ext' + · show _ = CommRingCat.ofHom C ≫ _ + rw [CommRingCat.comp_eq_ring_hom_comp, RingHom.comp_assoc, eval₂Hom_comp_C, + ← CommRingCat.comp_eq_ring_hom_comp, ← cancel_mono (Scheme.ΓSpecIso _).hom] + erw [← Scheme.comp_app] + rw [homOfVector_over, Scheme.comp_app] + simp only [TopologicalSpace.Opens.map_top, Category.assoc, Scheme.ΓSpecIso_naturality, ← + Scheme.toSpecΓ_app_top] + erw [← Scheme.comp_app_assoc] + rw [Scheme.isoSpec, asIso_inv, IsIso.hom_inv_id] + simp + rfl + · intro i + erw [CommRingCat.comp_apply, coe_eval₂Hom] + simp only [eval₂_X] + exact homOfVector_app_top_coord _ _ _ + +@[simp] +lemma isoOfIsAffine_hom_app_top [IsAffine S] : + (isoOfIsAffine n S).hom.app ⊤ = + (Scheme.ΓSpecIso _).hom ≫ eval₂Hom ((𝔸(n; S) ↘ S).app ⊤) (coord S) := by + simp [isoOfIsAffine_hom] + +@[local simp] +lemma isoOfIsAffine_inv_app_top_coord [IsAffine S] (i) : + (isoOfIsAffine n S).inv.app ⊤ (coord _ i) = (Scheme.ΓSpecIso (.of _)).inv (.X i) := + homOfVector_app_top_coord _ _ _ + +@[reassoc (attr := simp)] +lemma isoOfIsAffine_inv_over [IsAffine S] : + (isoOfIsAffine n S).inv ≫ 𝔸(n; S) ↘ S = Spec.map C ≫ S.isoSpec.inv := + pullback.lift_fst _ _ _ + +instance [IsAffine S] : IsAffine 𝔸(n; S) := isAffine_of_isIso (isoOfIsAffine n S).hom + +variable (n) in +/-- The affine space over an affine base is isomorphic to the spectrum of the polynomial ring. -/ +def SpecIso (R : CommRingCat.{max u v}) : + 𝔸(n; Spec R) ≅ Spec (.of (MvPolynomial n R)) := + isoOfIsAffine _ _ ≪≫ Scheme.Spec.mapIso (MvPolynomial.mapEquiv _ + (Scheme.ΓSpecIso R).symm.commRingCatIsoToRingEquiv).toCommRingCatIso.op + +@[simp] +lemma SpecIso_hom_app_top (R : CommRingCat.{max u v}) : + (SpecIso n R).hom.app ⊤ = (Scheme.ΓSpecIso _).hom ≫ + eval₂Hom ((Scheme.ΓSpecIso _).inv ≫ (𝔸(n; Spec R) ↘ Spec R).app ⊤) (coord (Spec R)) := by + simp only [SpecIso, Iso.trans_hom, Functor.mapIso_hom, Iso.op_hom, RingEquiv.toCommRingCatIso_hom, + RingEquiv.toRingHom_eq_coe, Scheme.Spec_map, Quiver.Hom.unop_op, Scheme.comp_coeBase, + TopologicalSpace.Opens.map_comp_obj, TopologicalSpace.Opens.map_top, Scheme.comp_app, + isoOfIsAffine_hom_app_top] + erw [Scheme.ΓSpecIso_naturality_assoc] + congr 1 + apply ringHom_ext' + · ext; simp; rfl + · simp + +@[local simp] +lemma SpecIso_inv_app_top_coord (R : CommRingCat.{max u v}) (i) : + (SpecIso n R).inv.app ⊤ (coord _ i) = (Scheme.ΓSpecIso (.of _)).inv (.X i) := by + simp only [SpecIso, Iso.trans_inv, Functor.mapIso_inv, Iso.op_inv, RingEquiv.toCommRingCatIso_inv, + mapEquiv_symm, RingEquiv.toRingHom_eq_coe, Scheme.Spec_map, Quiver.Hom.unop_op, + Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, TopologicalSpace.Opens.map_top, + Scheme.comp_app, CommRingCat.comp_apply] + erw [isoOfIsAffine_inv_app_top_coord, ← CommRingCat.comp_apply] + rw [← Scheme.ΓSpecIso_inv_naturality] + erw [CommRingCat.comp_apply] + congr 1 + exact map_X _ _ + +@[reassoc (attr := simp)] +lemma SpecIso_inv_over (R : CommRingCat.{max u v}) : + (SpecIso n R).inv ≫ 𝔸(n; Spec R) ↘ Spec R = Spec.map C := by + simp only [SpecIso, Iso.trans_inv, Functor.mapIso_inv, Iso.op_inv, RingEquiv.toCommRingCatIso_inv, + mapEquiv_symm, RingEquiv.toRingHom_eq_coe, Scheme.Spec_map, Quiver.Hom.unop_op, Category.assoc, + isoOfIsAffine_inv_over, Scheme.isoSpec_Spec_inv, ← Spec.map_comp] + congr 1 + rw [Iso.inv_comp_eq] + ext + exact map_C _ _ + +section functorial + +variable (n) in +/-- `𝔸(n; S)` is functorial wrt `S`. -/ +def map {S T : Scheme.{max u v}} (f : S ⟶ T) : 𝔸(n; S) ⟶ 𝔸(n; T) := + homOfVector (𝔸(n; S) ↘ S ≫ f) (coord S) + +@[reassoc (attr := simp)] +lemma map_over {S T : Scheme.{max u v}} (f : S ⟶ T) : map n f ≫ 𝔸(n; T) ↘ T = 𝔸(n; S) ↘ S ≫ f := + pullback.lift_fst _ _ _ + +@[local simp] +lemma map_app_top_coord {S T : Scheme.{max u v}} (f : S ⟶ T) (i) : + (map n f).app ⊤ (coord T i) = coord S i := + homOfVector_app_top_coord _ _ _ + +@[simp] +lemma map_id : map n (𝟙 S) = 𝟙 𝔸(n; S) := by + ext1 <;> simp [-TopologicalSpace.Opens.map_top]; rfl + +@[reassoc, simp] +lemma map_comp {S S' S'' : Scheme} (f : S ⟶ S') (g : S' ⟶ S'') : + map n (f ≫ g) = map n f ≫ map n g := by + ext1 + · simp + · simp only [TopologicalSpace.Opens.map_top, Scheme.comp_coeBase, + TopologicalSpace.Opens.map_comp_obj, Scheme.comp_app, CommRingCat.comp_apply] + erw [map_app_top_coord, map_app_top_coord, map_app_top_coord] + +lemma map_Spec_map {R S : CommRingCat.{max u v}} (φ : R ⟶ S) : + map n (Spec.map φ) = + (SpecIso n S).hom ≫ Spec.map (MvPolynomial.map φ) ≫ (SpecIso n R).inv := by + rw [← Iso.inv_comp_eq] + ext1 + · simp only [map_over, Category.assoc, SpecIso_inv_over, SpecIso_inv_over_assoc, + ← Spec.map_comp, CommRingCat.comp_eq_ring_hom_comp] + rw [map_comp_C] + · simp only [Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, + TopologicalSpace.Opens.map_top, Scheme.comp_app, CommRingCat.comp_apply] + conv_lhs => enter[2]; tactic => exact map_app_top_coord _ _ + conv_rhs => enter[2]; tactic => exact SpecIso_inv_app_top_coord _ _ + erw [SpecIso_inv_app_top_coord, ← CommRingCat.comp_apply] + rw [← Scheme.ΓSpecIso_inv_naturality] + erw [CommRingCat.comp_apply, map_X] + rfl + +/-- The map between affine spaces over affine bases is +isomorphic to the natural map between polynomial rings. -/ +def mapSpecMap {R S : CommRingCat.{max u v}} (φ : R ⟶ S) : + Arrow.mk (map n (Spec.map φ)) ≅ + Arrow.mk (Spec.map (CommRingCat.ofHom (MvPolynomial.map (σ := n) φ))) := + Arrow.isoMk (SpecIso n S) (SpecIso n R) (by simp [map_Spec_map]; rfl) + +/-- `𝔸(n; S)` is functorial wrt `n`. -/ +def reindex {n m : Type v} (i : m → n) (S : Scheme.{max u v}) : 𝔸(n; S) ⟶ 𝔸(m; S) := + homOfVector (𝔸(n; S) ↘ S) (coord S ∘ i) + +@[reassoc (attr := simp)] +lemma reindex_over {n m : Type v} (i : m → n) (S : Scheme.{max u v}) : + reindex i S ≫ 𝔸(m; S) ↘ S = 𝔸(n; S) ↘ S := + pullback.lift_fst _ _ _ + +@[local simp] +lemma reindex_app_top_coord {n m : Type v} (i : m → n) (S : Scheme.{max u v}) (j : m) : + (reindex i S).app ⊤ (coord S j) = coord S (i j) := + homOfVector_app_top_coord _ _ _ + +@[simp] +lemma reindex_id : reindex id S = 𝟙 𝔸(n; S) := by + ext1 <;> simp [-TopologicalSpace.Opens.map_top]; rfl + +lemma reindex_comp {n₁ n₂ n₃ : Type v} (i : n₁ → n₂) (j : n₂ → n₃) (S : Scheme.{max u v}) : + reindex (j ∘ i) S = reindex j S ≫ reindex i S := by + have H₁ : reindex (j ∘ i) S ≫ 𝔸(n₁; S) ↘ S = (reindex j S ≫ reindex i S) ≫ 𝔸(n₁; S) ↘ S := by simp + have H₂ (k) : (reindex (j ∘ i) S).app ⊤ (coord S k) = + (reindex j S).app ⊤ ((reindex i S).app ⊤ (coord S k)) := by + rw [reindex_app_top_coord, reindex_app_top_coord, reindex_app_top_coord] + rfl + exact hom_ext H₁ H₂ + +-- These time out if added to `reindex_comp` directly. +attribute [reassoc] reindex_comp +attribute [simp] reindex_comp + +lemma map_reindex {n₁ n₂ : Type v} (i : n₁ → n₂) {S T : Scheme.{max u v}} (f : S ⟶ T) : + map n₂ f ≫ reindex i T = reindex i S ≫ map n₁ f := by + apply hom_ext + · simp + · intro j + simp only [Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, Scheme.comp_app, + CommRingCat.comp_apply, map_app_top_coord, reindex_app_top_coord] + simp only [TopologicalSpace.Opens.map_top] + erw [map_app_top_coord f (i j), reindex_app_top_coord i S] + +attribute [reassoc (attr := simp)] map_reindex + +/-- The affine space as a functor. -/ +@[simps] +def functor : (Type v)ᵒᵖ ⥤ Scheme.{max u v} ⥤ Scheme.{max u v} where + obj n := { obj := AffineSpace n.unop, map := map n.unop, map_id := map_id, map_comp := map_comp } + map {n m} i := { app := reindex i.unop, naturality := fun _ _ ↦ map_reindex i.unop } + map_id n := by ext: 2; exact reindex_id _ + map_comp f g := by ext: 2; dsimp; exact reindex_comp _ _ _ + +end functorial + +end AffineSpace + +end AlgebraicGeometry From 8db27b2c62ba9902c27e8c309d875166d5505db8 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 16 Nov 2024 08:05:13 +0000 Subject: [PATCH 008/829] chore(MeasureTheory): rename `AbsolutelyContinuous.smul*` (#19044) I forgot to push it to #18953 before marking it ready for review. Rename `smul` to `smul_left` and `smul_both` to `smul`. This way it matches, e.g., `add*` lemmas. --- Mathlib/Analysis/Convex/Integral.lean | 5 ++--- .../Covering/Differentiation.lean | 10 ++++------ Mathlib/MeasureTheory/Measure/Haar/Unique.lean | 2 +- .../MeasureTheory/Measure/MeasureSpace.lean | 18 +++++++++++------- .../Measure/MutuallySingular.lean | 2 +- 5 files changed, 19 insertions(+), 18 deletions(-) diff --git a/Mathlib/Analysis/Convex/Integral.lean b/Mathlib/Analysis/Convex/Integral.lean index 7d933906dae1b..eb9be5b719ffe 100644 --- a/Mathlib/Analysis/Convex/Integral.lean +++ b/Mathlib/Analysis/Convex/Integral.lean @@ -81,9 +81,8 @@ theorem Convex.integral_mem [IsProbabilityMeasure μ] (hs : Convex ℝ s) (hsc : integrable function sending `μ`-a.e. points to `s`, then the average value of `f` belongs to `s`: `⨍ x, f x ∂μ ∈ s`. See also `Convex.centerMass_mem` for a finite sum version of this lemma. -/ theorem Convex.average_mem [IsFiniteMeasure μ] [NeZero μ] (hs : Convex ℝ s) (hsc : IsClosed s) - (hfs : ∀ᵐ x ∂μ, f x ∈ s) (hfi : Integrable f μ) : (⨍ x, f x ∂μ) ∈ s := by - refine hs.integral_mem hsc (ae_mono' ?_ hfs) hfi.to_average - exact AbsolutelyContinuous.smul (refl _) _ + (hfs : ∀ᵐ x ∂μ, f x ∈ s) (hfi : Integrable f μ) : (⨍ x, f x ∂μ) ∈ s := + hs.integral_mem hsc (ae_mono' smul_absolutelyContinuous hfs) hfi.to_average /-- If `μ` is a non-zero finite measure on `α`, `s` is a convex closed set in `E`, and `f` is an integrable function sending `μ`-a.e. points to `s`, then the average value of `f` belongs to `s`: diff --git a/Mathlib/MeasureTheory/Covering/Differentiation.lean b/Mathlib/MeasureTheory/Covering/Differentiation.lean index f2de314db2a9e..eb17b94c41db8 100644 --- a/Mathlib/MeasureTheory/Covering/Differentiation.lean +++ b/Mathlib/MeasureTheory/Covering/Differentiation.lean @@ -178,7 +178,7 @@ theorem ae_eventually_measure_zero_of_singular (hρ : ρ ⟂ₘ μ) : rw [ENNReal.mul_inv_cancel (ENNReal.coe_pos.2 εpos).ne' ENNReal.coe_ne_top, one_mul] _ ≤ (ε : ℝ≥0∞)⁻¹ * ρ (s ∩ o) := by gcongr - refine v.measure_le_of_frequently_le ρ ((Measure.AbsolutelyContinuous.refl μ).smul ε) _ ?_ + refine v.measure_le_of_frequently_le ρ smul_absolutelyContinuous _ ?_ intro x hx rw [hs] at hx simp only [mem_inter_iff, not_lt, not_eventually, mem_setOf_eq] at hx @@ -224,9 +224,7 @@ theorem null_of_frequently_le_of_frequently_ge {c d : ℝ≥0} (hcd : c < d) (s _ < d * μ s' := by apply (ENNReal.mul_lt_mul_right h _).2 (ENNReal.coe_lt_coe.2 hcd) exact (lt_of_le_of_lt (measure_mono inter_subset_right) μo).ne - _ ≤ ρ s' := - v.measure_le_of_frequently_le ρ ((Measure.AbsolutelyContinuous.refl μ).smul d) s' fun x hx => - hd x hx.1 + _ ≤ ρ s' := v.measure_le_of_frequently_le ρ smul_absolutelyContinuous s' fun x hx ↦ hd x hx.1 /-- If `ρ` is absolutely continuous with respect to `μ`, then for almost every `x`, the ratio `ρ a / μ a` converges as `a` shrinks to `x` along a Vitali family for `μ`. -/ @@ -383,7 +381,7 @@ theorem exists_measurable_supersets_limRatio {p q : ℝ≥0} (hpq : p < q) : rw [inter_comm, measure_toMeasurable_add_inter_right (measurableSet_toMeasurable _ _) J] _ ≤ ρ (toMeasurable (ρ + μ) (u m) ∩ w n) := by rw [← coe_nnreal_smul_apply] - refine v.measure_le_of_frequently_le _ (AbsolutelyContinuous.rfl.smul _) _ ?_ + refine v.measure_le_of_frequently_le _ (.smul_left .rfl _) _ ?_ intro x hx have L : Tendsto (fun a : Set α => ρ a / μ a) (v.filterAt x) (𝓝 (v.limRatio ρ x)) := tendsto_nhds_limUnder hx.2.1.1 @@ -464,7 +462,7 @@ theorem mul_measure_le_of_subset_lt_limRatioMeas {q : ℝ≥0} {s : Set α} _ ≤ ρ (s ∩ t) + (q • μ) tᶜ := by gcongr; apply inter_subset_right _ = ρ (s ∩ t) := by simp [A] _ ≤ ρ s := by gcongr; apply inter_subset_left - refine v.measure_le_of_frequently_le _ (AbsolutelyContinuous.rfl.smul _) _ ?_ + refine v.measure_le_of_frequently_le _ (.smul_left .rfl _) _ ?_ intro x hx have I : ∀ᶠ a in v.filterAt x, (q : ℝ≥0∞) < ρ a / μ a := (tendsto_order.1 hx.2).1 _ (h hx.1) apply I.frequently.mono fun a ha => ?_ diff --git a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean index fb3d94090f46d..a8474fe9a8a32 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean @@ -910,7 +910,7 @@ theorem absolutelyContinuous_isHaarMeasure [LocallyCompactSpace G] have h : haarMeasure K = (haarScalarFactor (haarMeasure K) ν : ℝ≥0∞) • ν := isMulLeftInvariant_eq_smul (haarMeasure K) ν rw [haarMeasure_unique μ K, h, smul_smul] - exact AbsolutelyContinuous.smul (Eq.absolutelyContinuous rfl) _ + exact smul_absolutelyContinuous /-- A continuous surjective monoid homomorphism of topological groups with compact codomain is measure preserving, provided that the Haar measures on the domain and on the codomain diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index ea2750422043c..e80720475f839 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -1521,16 +1521,21 @@ protected theorem trans (h1 : μ₁ ≪ μ₂) (h2 : μ₂ ≪ μ₃) : μ₁ protected theorem map (h : μ ≪ ν) {f : α → β} (hf : Measurable f) : μ.map f ≪ ν.map f := AbsolutelyContinuous.mk fun s hs => by simpa [hf, hs] using @h _ -protected theorem smul [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] (h : μ ≪ ν) (c : R) : +protected theorem smul_left [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] (h : μ ≪ ν) (c : R) : c • μ ≪ ν := fun s hνs => by simp only [h hνs, smul_apply, smul_zero, ← smul_one_smul ℝ≥0∞ c (0 : ℝ≥0∞)] -protected theorem smul_both [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] (h : μ ≪ ν) (c : R) : +/-- If `μ ≪ ν`, then `c • μ ≪ c • ν`. + +Earlier, this name was used for what's now called `AbsolutelyContinuous.smul_left`. -/ +protected theorem smul [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] (h : μ ≪ ν) (c : R) : c • μ ≪ c • ν := by intro s hνs rw [smul_apply, ← smul_one_smul ℝ≥0∞, smul_eq_mul, mul_eq_zero] at hνs ⊢ exact hνs.imp_right fun hs ↦ h hs +@[deprecated (since := "2024-11-14")] protected alias smul_both := AbsolutelyContinuous.smul + protected lemma add (h1 : μ₁ ≪ ν) (h2 : μ₂ ≪ ν') : μ₁ + μ₂ ≪ ν + ν' := by intro s hs simp only [coe_add, Pi.add_apply, add_eq_zero] at hs ⊢ @@ -1541,9 +1546,8 @@ lemma add_left_iff {μ₁ μ₂ ν : Measure α} : refine ⟨fun h ↦ ?_, fun h ↦ (h.1.add h.2).trans ?_⟩ · have : ∀ s, ν s = 0 → μ₁ s = 0 ∧ μ₂ s = 0 := by intro s hs0; simpa using h hs0 exact ⟨fun s hs0 ↦ (this s hs0).1, fun s hs0 ↦ (this s hs0).2⟩ - · have : ν + ν = 2 • ν := by ext; simp [two_mul] - rw [this] - exact AbsolutelyContinuous.rfl.smul 2 + · rw [← two_smul ℝ≥0] + exact AbsolutelyContinuous.rfl.smul_left 2 lemma add_left {μ₁ μ₂ ν : Measure α} (h₁ : μ₁ ≪ ν) (h₂ : μ₂ ≪ ν) : μ₁ + μ₂ ≪ ν := Measure.AbsolutelyContinuous.add_left_iff.mpr ⟨h₁, h₂⟩ @@ -1572,7 +1576,7 @@ lemma absolutelyContinuous_sum_right {μs : ι → Measure α} (i : ι) (hνμ : simp only [sum_apply _ hs, ENNReal.tsum_eq_zero] at hs0 exact hνμ (hs0 i) -lemma smul_absolutelyContinuous {c : ℝ≥0∞} : c • μ ≪ μ := .smul .rfl _ +lemma smul_absolutelyContinuous {c : ℝ≥0∞} : c • μ ≪ μ := .smul_left .rfl _ theorem absolutelyContinuous_of_le_smul {μ' : Measure α} {c : ℝ≥0∞} (hμ'_le : μ' ≤ c • μ) : μ' ≪ μ := @@ -1655,7 +1659,7 @@ protected theorem aemeasurable (hf : QuasiMeasurePreserving f μa μb) : AEMeasu theorem smul_measure {R : Type*} [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ ℝ≥0∞] (hf : QuasiMeasurePreserving f μa μb) (c : R) : QuasiMeasurePreserving f (c • μa) (c • μb) := - ⟨hf.1, by rw [Measure.map_smul]; exact hf.2.smul_both c⟩ + ⟨hf.1, by rw [Measure.map_smul]; exact hf.2.smul c⟩ theorem ae_map_le (h : QuasiMeasurePreserving f μa μb) : ae (μa.map f) ≤ ae μb := h.2.ae_le diff --git a/Mathlib/MeasureTheory/Measure/MutuallySingular.lean b/Mathlib/MeasureTheory/Measure/MutuallySingular.lean index 4cb10bc11a5f0..257f1ecc3765f 100644 --- a/Mathlib/MeasureTheory/Measure/MutuallySingular.lean +++ b/Mathlib/MeasureTheory/Measure/MutuallySingular.lean @@ -128,7 +128,7 @@ theorem add_right (h₁ : μ ⟂ₘ ν₁) (h₂ : μ ⟂ₘ ν₂) : μ ⟂ₘ add_right_iff.2 ⟨h₁, h₂⟩ theorem smul (r : ℝ≥0∞) (h : ν ⟂ₘ μ) : r • ν ⟂ₘ μ := - h.mono_ac (AbsolutelyContinuous.rfl.smul r) AbsolutelyContinuous.rfl + h.mono_ac (AbsolutelyContinuous.rfl.smul_left r) AbsolutelyContinuous.rfl theorem smul_nnreal (r : ℝ≥0) (h : ν ⟂ₘ μ) : r • ν ⟂ₘ μ := h.smul r From f136306a91e817e991007c693d16a53cd824256c Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Sat, 16 Nov 2024 08:23:24 +0000 Subject: [PATCH 009/829] feat(Archive/Imo): IMO 2024 Q5 (#18685) Formalize IMO 2024 problem 5. Note that this solution is longer than any of the existing IMO formalizations in either the mathlib archive or compfiles; feel free to golf, where doing so isn't adverse to the role of such formalizations as readable examples. Various definitions using proofs (that certain paths used in an optimal strategy are valid paths) by essentially `split_ifs <;> simp <;> omega` are (subjectively, not counted heartbeats) slow for Lean to process; in those cases, I think the tactics used are a better approach as an example than a faster but lower-level and probably longer proof would be, so rather than making those proofs longer to speed them up it might be better to treat them as examples that might show opportunities for speeding up those tactics. --- Archive.lean | 1 + Archive/Imo/Imo2024Q5.lean | 1046 ++++++++++++++++++++++++++++++++++++ 2 files changed, 1047 insertions(+) create mode 100644 Archive/Imo/Imo2024Q5.lean diff --git a/Archive.lean b/Archive.lean index 97425492077d2..3b1428fa31cb4 100644 --- a/Archive.lean +++ b/Archive.lean @@ -45,6 +45,7 @@ import Archive.Imo.Imo2020Q2 import Archive.Imo.Imo2021Q1 import Archive.Imo.Imo2024Q1 import Archive.Imo.Imo2024Q2 +import Archive.Imo.Imo2024Q5 import Archive.Imo.Imo2024Q6 import Archive.MiuLanguage.Basic import Archive.MiuLanguage.DecisionNec diff --git a/Archive/Imo/Imo2024Q5.lean b/Archive/Imo/Imo2024Q5.lean new file mode 100644 index 0000000000000..7b090f48cc987 --- /dev/null +++ b/Archive/Imo/Imo2024Q5.lean @@ -0,0 +1,1046 @@ +/- +Copyright (c) 2024 Joseph Myers. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joseph Myers +-/ +import Mathlib.Data.Fin.VecNotation +import Mathlib.Data.List.ChainOfFn +import Mathlib.Data.Nat.Dist +import Mathlib.Order.Fin.Basic +import Mathlib.Tactic.IntervalCases + +/-! +# IMO 2024 Q5 + +Turbo the snail plays a game on a board with $2024$ rows and $2023$ columns. There are hidden +monsters in $2022$ of the cells. Initially, Turbo does not know where any of the monsters are, +but he knows that there is exactly one monster in each row except the first row and the last +row, and that each column contains at most one monster. + +Turbo makes a series of attempts to go from the first row to the last row. On each attempt, +he chooses to start on any cell in the first row, then repeatedly moves to an adjacent cell +sharing a common side. (He is allowed to return to a previously visited cell.) If he reaches a +cell with a monster, his attempt ends and he is transported back to the first row to start a +new attempt. The monsters do not move, and Turbo remembers whether or not each cell he has +visited contains a monster. If he reaches any cell in the last row, his attempt ends and the +game is over. + +Determine the minimum value of $n$ for which Turbo has a strategy that guarantees reaching +the last row on the $n$th attempt or earlier, regardless of the locations of the monsters. + +We follow the solution from the +[official solutions](https://www.imo2024.uk/s/IMO2024-solutions-updated.pdf). To show that $n$ +is at least $3$, it is possible that the first cell Turbo encounters in the second row on his +first attempt contains a monster, and also possible that the first cell Turbo encounters in the +third row on his second attempt contains a monster. To show that $3$ attempts suffice, the first +attempt can be used to locate the monster in the second row; if this is not at either side of +the board, two more attempts suffice to pass behind that monster and from there go along its +column to the last row, while if it is at one side of the board, the second attempt follows a +zigzag path such that if it encounters a monster the third attempt can avoid all monsters. +-/ + + +namespace Imo2024Q5 + +/-! ### Definitions for setting up the problem -/ + +-- There are N monsters, N+1 columns and N+2 rows. +variable {N : ℕ} + +/-- A cell on the board for the game. -/ +abbrev Cell (N : ℕ) : Type := Fin (N + 2) × Fin (N + 1) + +/-- A row that is neither the first nor the last (and thus contains a monster). -/ +abbrev InteriorRow (N : ℕ) : Type := (Set.Icc 1 ⟨N, by omega⟩ : Set (Fin (N + 2))) + +/-- Data for valid positions of the monsters. -/ +abbrev MonsterData (N : ℕ) : Type := InteriorRow N ↪ Fin (N + 1) + +/-- The cells with monsters as a Set, given an injection from rows to columns. -/ +def MonsterData.monsterCells (m : MonsterData N) : + Set (Cell N) := + Set.range (fun x : InteriorRow N ↦ ((x : Fin (N + 2)), m x)) + +/-- Whether two cells are adjacent. -/ +def Adjacent (x y : Cell N) : Prop := + Nat.dist x.1 y.1 + Nat.dist x.2 y.2 = 1 + +/-- A valid path from the first to the last row. -/ +structure Path (N : ℕ) where + /-- The cells on the path. -/ + cells : List (Cell N) + nonempty : cells ≠ [] + head_first_row : (cells.head nonempty).1 = 0 + last_last_row : (cells.getLast nonempty).1 = N + 1 + valid_move_seq : cells.Chain' Adjacent + +/-- The first monster on a path, or `none`. -/ +noncomputable def Path.firstMonster (p : Path N) (m : MonsterData N) : Option (Cell N) := + letI := Classical.propDecidable + p.cells.find? (fun x ↦ (x ∈ m.monsterCells : Bool)) + +/-- A strategy, given the results of initial attempts, returns a path for the next attempt. -/ +abbrev Strategy (N : ℕ) : Type := ⦃k : ℕ⦄ → (Fin k → Option (Cell N)) → Path N + +/-- Playing a strategy, k attempts. -/ +noncomputable def Strategy.play (s : Strategy N) (m : MonsterData N) : + (k : ℕ) → Fin k → Option (Cell N) +| 0 => Fin.elim0 +| k + 1 => Fin.snoc (s.play m k) ((s (s.play m k)).firstMonster m) + +/-- The predicate for a strategy winning within the given number of attempts. -/ +def Strategy.WinsIn (s : Strategy N) (m : MonsterData N) (k : ℕ) : Prop := + none ∈ Set.range (s.play m k) + +/-- Whether a strategy forces a win within the given number of attempts. -/ +def Strategy.ForcesWinIn (s : Strategy N) (k : ℕ) : Prop := + ∀ m, s.WinsIn m k + +/-! ### API definitions and lemmas about `Cell` -/ + +/-- Reflecting a cell of the board (swapping left and right sides of the board). -/ +def Cell.reflect (c : Cell N) : Cell N := (c.1, c.2.rev) + +/-! ### API definitions and lemmas about `MonsterData` -/ + +/-- The row 1, in the form required for MonsterData. -/ +def row1 (hN : 2 ≤ N) : InteriorRow N := + ⟨1, ⟨by omega, by + rw [Fin.le_def] + simp + omega⟩⟩ + +lemma coe_coe_row1 (hN : 2 ≤ N) : (((row1 hN) : Fin (N + 2)) : ℕ) = 1 := + rfl + +/-- The row 2, in the form required for MonsterData. -/ +def row2 (hN : 2 ≤ N) : InteriorRow N := + ⟨⟨2, by omega⟩, ⟨by + simp only [Fin.le_def, Fin.val_one] + omega, by + simp only [Fin.le_def] + omega⟩⟩ + +/-- Reflecting monster data. -/ +def MonsterData.reflect (m : MonsterData N) : MonsterData N where + toFun := Fin.rev ∘ m + inj' := fun i j hij ↦ by simpa using hij + +lemma MonsterData.reflect_reflect (m : MonsterData N) : m.reflect.reflect = m := by + ext i + simp [MonsterData.reflect] + +lemma MonsterData.not_mem_monsterCells_of_fst_eq_zero (m : MonsterData N) + {c : Cell N} (hc : c.1 = 0) : c ∉ m.monsterCells := by + simp [monsterCells, Prod.ext_iff, hc] + +lemma MonsterData.le_N_of_mem_monsterCells {m : MonsterData N} {c : Cell N} + (hc : c ∈ m.monsterCells) : (c.1 : ℕ) ≤ N := by + simp only [monsterCells, Set.mem_range, Subtype.exists, Set.mem_Icc] at hc + rcases hc with ⟨r, ⟨h1, hN⟩, rfl⟩ + rw [Fin.le_def] at hN + exact hN + +lemma MonsterData.mk_mem_monsterCells_iff_of_le {m : MonsterData N} {r : Fin (N + 2)} + (hr1 : 1 ≤ r) (hrN : r ≤ ⟨N, by omega⟩) {c : Fin (N + 1)} : + (r, c) ∈ m.monsterCells ↔ m ⟨r, hr1, hrN⟩ = c := by + simp only [monsterCells, Set.mem_range, Prod.mk.injEq] + refine ⟨?_, ?_⟩ + · rintro ⟨r', rfl, rfl⟩ + rfl + · rintro rfl + exact ⟨⟨r, hr1, hrN⟩, rfl, rfl⟩ + +lemma MonsterData.mem_monsterCells_iff_of_le {m : MonsterData N} {x : Cell N} + (hr1 : 1 ≤ x.1) (hrN : x.1 ≤ ⟨N, by omega⟩) : + x ∈ m.monsterCells ↔ m ⟨x.1, hr1, hrN⟩ = x.2 := + MonsterData.mk_mem_monsterCells_iff_of_le hr1 hrN + +lemma MonsterData.mk_mem_monsterCells_iff {m : MonsterData N} {r : Fin (N + 2)} + {c : Fin (N + 1)} : + (r, c) ∈ m.monsterCells ↔ ∃ (hr1 : 1 ≤ r) (hrN : r ≤ ⟨N, by omega⟩), m ⟨r, hr1, hrN⟩ = c := by + refine ⟨fun h ↦ ?_, fun ⟨hr1, hrN, h⟩ ↦ (mem_monsterCells_iff_of_le hr1 hrN).2 h⟩ + rcases h with ⟨⟨mr, hr1, hrN⟩, h⟩ + simp only [Prod.mk.injEq] at h + rcases h with ⟨rfl, rfl⟩ + exact ⟨hr1, hrN, rfl⟩ + +/-! ### API definitions and lemmas about `Adjacent` -/ + +lemma Adjacent.le_of_lt {x y : Cell N} (ha : Adjacent x y) {r : Fin (N + 2)} (hr : r < y.1) : + r ≤ x.1 := by + rw [Adjacent, Nat.dist] at ha + omega + +/-! ### API definitions and lemmas about `Path` -/ + +lemma Path.exists_mem_fst_eq (p : Path N) (r : Fin (N + 2)) : ∃ c ∈ p.cells, c.1 = r := by + let i : ℕ := p.cells.findIdx fun c ↦ r ≤ c.1 + have hi : i < p.cells.length := by + refine List.findIdx_lt_length_of_exists ⟨p.cells.getLast p.nonempty, ?_⟩ + simp only [List.getLast_mem, p.last_last_row, decide_eq_true_eq, true_and] + rw [Fin.le_def, Fin.add_def] + have h := r.isLt + rw [Nat.lt_succ_iff] at h + convert h + simp + have hig : r ≤ (p.cells[i]).1 := of_decide_eq_true (List.findIdx_getElem (w := hi)) + refine ⟨p.cells[i], List.getElem_mem _, ?_⟩ + rcases hig.lt_or_eq.symm with h | h + · exact h.symm + · rcases Nat.eq_zero_or_pos i with hi | hi + · simp only [hi, List.getElem_zero, p.head_first_row, Fin.not_lt_zero] at h + · suffices r ≤ p.cells[i - 1].1 by + exfalso + have hi' : i - 1 < i := by omega + exact of_decide_eq_false (List.not_of_lt_findIdx hi') this + have ha : Adjacent p.cells[i - 1] p.cells[i] := by + convert List.chain'_iff_get.1 p.valid_move_seq (i - 1) ?_ + · simp [Nat.sub_add_cancel hi] + · omega + exact ha.le_of_lt h + +lemma Path.exists_mem_le_fst (p : Path N) (r : Fin (N + 2)) : ∃ c ∈ p.cells, r ≤ c.1 := by + rcases p.exists_mem_fst_eq r with ⟨c, hc, he⟩ + exact ⟨c, hc, he.symm.le⟩ + +/-- The first path element on a row. -/ +def Path.findFstEq (p : Path N) (r : Fin (N + 2)) : Cell N := + (p.cells.find? (fun c ↦ c.1 = r)).get (List.find?_isSome.2 (by simpa using p.exists_mem_fst_eq r)) + +lemma Path.find_eq_some_findFstEq (p : Path N) (r : Fin (N + 2)) : + p.cells.find? (fun c ↦ c.1 = r) = some (p.findFstEq r) := by + rw [Option.eq_some_iff_get_eq] + exact ⟨_, rfl⟩ + +lemma Path.findFstEq_mem_cells (p : Path N) (r : Fin (N + 2)) : p.findFstEq r ∈ p.cells := + List.mem_of_find?_eq_some (p.find_eq_some_findFstEq r) + +lemma Path.findFstEq_fst (p : Path N) (r : Fin (N + 2)) : (p.findFstEq r).1 = r := by + have h := List.find?_some (p.find_eq_some_findFstEq r) + simpa using h + +lemma find?_eq_eq_find?_le {l : List (Cell N)} {r : Fin (N + 2)} (hne : l ≠ []) + (hf : (l.head hne).1 ≤ r) (ha : l.Chain' Adjacent) : + l.find? (fun c ↦ c.1 = r) = l.find? (fun c ↦ r ≤ c.1) := by + induction l + case nil => simp at hne + case cons head tail hi => + by_cases h : head.1 = r + · simp [h] + · have h' : ¬(r ≤ head.1) := fun hr' ↦ h (le_antisymm hf hr') + simp only [h, decide_False, Bool.false_eq_true, not_false_eq_true, List.find?_cons_of_neg, h'] + rcases tail with ⟨⟩ | ⟨htail, ttail⟩ + · simp + · simp only [List.chain'_cons] at ha + rcases ha with ⟨ha1, ha2⟩ + simp only [List.head_cons] at hf + simp only [ne_eq, reduceCtorEq, not_false_eq_true, List.head_cons, ha2, true_implies] at hi + refine hi ?_ + simp only [Adjacent, Nat.dist] at ha1 + omega + +lemma Path.findFstEq_eq_find?_le (p : Path N) (r : Fin (N + 2)) : p.findFstEq r = + (p.cells.find? (fun c ↦ r ≤ c.1)).get + (List.find?_isSome.2 (by simpa using p.exists_mem_le_fst r)) := by + rw [findFstEq] + convert rfl using 2 + refine (find?_eq_eq_find?_le p.nonempty ?_ p.valid_move_seq).symm + simp [p.head_first_row] + +lemma Path.firstMonster_isSome {p : Path N} {m : MonsterData N} : + (p.firstMonster m).isSome = true ↔ ∃ x, x ∈ p.cells ∧ x ∈ m.monsterCells := by + convert List.find?_isSome + simp + +lemma Path.firstMonster_eq_none {p : Path N} {m : MonsterData N} : + (p.firstMonster m) = none ↔ ∀ x, x ∈ p.cells → x ∉ m.monsterCells := by + convert List.find?_eq_none + simp + +lemma Path.one_lt_length_cells (p : Path N): 1 < p.cells.length := by + by_contra hl + have h : p.cells.length = 0 ∨ p.cells.length = 1 := by omega + rcases h with h | h + · rw [List.length_eq_zero] at h + exact p.nonempty h + · rw [List.length_eq_one] at h + rcases h with ⟨c, hc⟩ + have h1 := p.head_first_row + simp_rw [hc, List.head_cons] at h1 + have h2 := p.last_last_row + simp [hc, List.getLast_singleton, h1, Fin.add_def, Fin.ext_iff] at h2 + +/-- Remove the first cell from a path, if the second cell is also on the first row. -/ +def Path.tail (p : Path N) : Path N where + cells := if (p.cells[1]'p.one_lt_length_cells).1 = 0 then p.cells.tail else p.cells + nonempty := by + split_ifs + · rw [← List.length_pos_iff_ne_nil, List.length_tail, Nat.sub_pos_iff_lt] + exact p.one_lt_length_cells + · exact p.nonempty + head_first_row := by + split_ifs with h + · convert h + rw [List.head_tail] + · exact p.head_first_row + last_last_row := by + split_ifs + · rw [← p.last_last_row, List.getLast_tail _ _] + · exact p.last_last_row + valid_move_seq := by + split_ifs + · exact List.Chain'.tail p.valid_move_seq + · exact p.valid_move_seq + +lemma Path.tail_induction {motive : Path N → Prop} (ind : ∀ p, motive p.tail → motive p) + (base : ∀ p, (p.cells[1]'p.one_lt_length_cells).1 ≠ 0 → motive p) (p : Path N) : motive p := by + rcases p with ⟨cells, nonempty, head_first_row, last_last_row, valid_move_seq⟩ + let p' : Path N := ⟨cells, nonempty, head_first_row, last_last_row, valid_move_seq⟩ + induction cells + case nil => simp at nonempty + case cons head tail hi => + by_cases h : (p'.cells[1]'p'.one_lt_length_cells).1 = 0 + · refine ind p' ?_ + simp_rw [Path.tail, if_pos h, List.tail_cons] + exact hi _ _ _ _ + · exact base p' h + +lemma Path.tail_findFstEq (p : Path N) {r : Fin (N + 2)} (hr : r ≠ 0) : + p.tail.findFstEq r = p.findFstEq r := by + by_cases h : (p.cells[1]'p.one_lt_length_cells).1 = 0 + · simp_rw [Path.tail, if_pos h] + nth_rw 2 [Path.findFstEq] + rcases p with ⟨cells, nonempty, head_first_row, last_last_row, valid_move_seq⟩ + rcases cells with ⟨⟩ | ⟨head, tail⟩ + · simp at nonempty + · simp only [List.head_cons] at head_first_row + simp only [List.find?_cons, head_first_row, hr.symm, decide_False] + rfl + · simp_rw [Path.tail, if_neg h] + +lemma Path.tail_firstMonster (p : Path N) (m : MonsterData N) : + p.tail.firstMonster m = p.firstMonster m := by + by_cases h : (p.cells[1]'p.one_lt_length_cells).1 = 0 + · simp_rw [Path.tail, if_pos h] + nth_rw 2 [Path.firstMonster] + rcases p with ⟨cells, nonempty, head_first_row, last_last_row, valid_move_seq⟩ + rcases cells with ⟨⟩ | ⟨head, tail⟩ + · simp at nonempty + · simp only [List.head_cons] at head_first_row + simp only [List.find?_cons, head_first_row, + m.not_mem_monsterCells_of_fst_eq_zero head_first_row, decide_False] + rfl + · simp_rw [Path.tail, if_neg h] + +lemma Path.firstMonster_eq_of_findFstEq_mem {p : Path N} {m : MonsterData N} + (h : p.findFstEq 1 ∈ m.monsterCells) : p.firstMonster m = some (p.findFstEq 1) := by + induction p using Path.tail_induction + case base p h0 => + have hl := p.one_lt_length_cells + have adj : Adjacent p.cells[0] p.cells[1] := List.chain'_iff_get.1 p.valid_move_seq 0 (by omega) + simp_rw [Adjacent, Nat.dist] at adj + have hc0 : (p.cells[0].1 : ℕ) = 0 := by + convert Fin.ext_iff.1 p.head_first_row + exact List.getElem_zero _ + have hc1 : (p.cells[1].1 : ℕ) ≠ 0 := Fin.val_ne_iff.2 h0 + have h1 : (p.cells[1].1 : ℕ) = 1 := by omega + simp_rw [firstMonster, findFstEq] + rcases p with ⟨cells, nonempty, head_first_row, last_last_row, valid_move_seq⟩ + rcases cells with ⟨⟩ | ⟨head, tail⟩ + · simp at nonempty + · simp only [List.head_cons] at head_first_row + simp only [List.getElem_cons_succ] at h1 + simp only [List.length_cons, lt_add_iff_pos_left, List.length_pos_iff_ne_nil] at hl + simp only [m.not_mem_monsterCells_of_fst_eq_zero head_first_row, decide_False, + Bool.false_eq_true, not_false_eq_true, List.find?_cons_of_neg, head_first_row, + Fin.zero_eq_one_iff, Nat.reduceEqDiff, Option.some_get] + simp only [findFstEq, head_first_row, Fin.zero_eq_one_iff, Nat.reduceEqDiff, decide_False, + Bool.false_eq_true, not_false_eq_true, List.find?_cons_of_neg] at h + rcases tail with ⟨⟩ | ⟨htail, ttail⟩ + · simp at hl + · simp only [List.getElem_cons_zero] at h1 + have h1' : htail.1 = 1 := by simp [Fin.ext_iff, h1] + simp only [h1', decide_True, List.find?_cons_of_pos, Option.get_some] at h + simp only [h1', h, decide_True, List.find?_cons_of_pos] + case ind p ht => + have h1 : (1 : Fin (N + 2)) ≠ 0 := by simp + rw [p.tail_findFstEq h1, p.tail_firstMonster m] at ht + exact ht h + +lemma Path.findFstEq_fst_sub_one_mem (p : Path N) {r : Fin (N + 2)} (hr : r ≠ 0) : + (⟨(r : ℕ) - 1, by omega⟩, (p.findFstEq r).2) ∈ p.cells := by + rw [findFstEq_eq_find?_le] + have h1 := p.one_lt_length_cells + obtain ⟨cr, hcrc, hcrr⟩ := p.exists_mem_fst_eq r + rcases p with ⟨cells, nonempty, head_first_row, last_last_row, valid_move_seq⟩ + dsimp only at h1 hcrc ⊢ + have hd : ∃ c ∈ cells, decide (r ≤ c.1) = true := ⟨cr, hcrc, by simpa using hcrr.symm.le⟩ + have hd' : cells.dropWhile (fun c ↦ ! decide (r ≤ c.1)) ≠ [] := by simpa using hd + have ht : cells.takeWhile (fun c ↦ ! decide (r ≤ c.1)) ≠ [] := by + intro h + rw [List.takeWhile_eq_nil_iff] at h + replace h := h (by omega) + simp [List.getElem_zero, head_first_row, hr] at h + simp_rw [cells.find?_eq_head_dropWhile_not hd, Option.get_some] + rw [← cells.takeWhile_append_dropWhile (fun c ↦ ! decide (r ≤ c.1)), + List.chain'_append] at valid_move_seq + have ha := valid_move_seq.2.2 + simp only [List.head?_eq_head hd', List.getLast?_eq_getLast _ ht, Option.mem_def, + Option.some.injEq, forall_eq'] at ha + nth_rw 1 [← cells.takeWhile_append_dropWhile (fun c ↦ ! decide (r ≤ c.1))] + refine List.mem_append_left _ ?_ + convert List.getLast_mem ht using 1 + have htr : ((List.takeWhile (fun c ↦ !decide (r ≤ c.1)) cells).getLast ht).1 < r := by + simpa using List.mem_takeWhile_imp (List.getLast_mem ht) + have hdr : r ≤ ((List.dropWhile (fun c ↦ !decide (r ≤ c.1)) cells).head hd').1 := by + simpa using cells.head_dropWhile_not (fun c ↦ !decide (r ≤ c.1)) hd' + simp only [Adjacent, Nat.dist] at ha + have hrm : N + 1 + (r : ℕ) = N + 2 + (r - 1) := by omega + simp only [Prod.ext_iff, Fin.ext_iff] + omega + +lemma Path.mem_of_firstMonster_eq_some {p : Path N} {m : MonsterData N} {c : Cell N} + (h : p.firstMonster m = some c) : c ∈ p.cells ∧ c ∈ m.monsterCells := by + simp_rw [firstMonster] at h + have h₁ := List.mem_of_find?_eq_some h + have h₂ := List.find?_some h + simp only [decide_eq_true_eq] at h₂ + exact ⟨h₁, h₂⟩ + +/-- Convert a function giving the cells of a path to a `Path`. -/ +def Path.ofFn {m : ℕ} (f : Fin m → Cell N) (hm : m ≠ 0) + (hf : (f ⟨0, Nat.pos_of_ne_zero hm⟩).1 = 0) + (hl : (f ⟨m - 1, Nat.sub_one_lt hm⟩).1 = ⟨N + 1, by omega⟩) + (ha : ∀ (i : ℕ) (hi : i + 1 < m), Adjacent (f ⟨i, Nat.lt_of_succ_lt hi⟩) (f ⟨i + 1, hi⟩)) : + Path N where + cells := List.ofFn f + nonempty := mt List.ofFn_eq_nil_iff.1 hm + head_first_row := by + rw [List.head_ofFn, hf] + last_last_row := by + simp [List.getLast_ofFn, hl, Fin.ext_iff, Fin.add_def] + valid_move_seq := by + rwa [List.chain'_ofFn] + +lemma Path.ofFn_cells {m : ℕ} (f : Fin m → Cell N) (hm : m ≠ 0) + (hf : (f ⟨0, Nat.pos_of_ne_zero hm⟩).1 = 0) + (hl : (f ⟨m - 1, Nat.sub_one_lt hm⟩).1 = ⟨N + 1, by omega⟩) + (ha : ∀ (i : ℕ) (hi : i + 1 < m), Adjacent (f ⟨i, Nat.lt_of_succ_lt hi⟩) (f ⟨i + 1, hi⟩)) : + (Path.ofFn f hm hf hl ha).cells = List.ofFn f := + rfl + +lemma Path.ofFn_firstMonster_eq_none {m : ℕ} (f : Fin m → Cell N) (hm hf hl ha) + (m : MonsterData N) : + ((Path.ofFn f hm hf hl ha).firstMonster m) = none ↔ ∀ i, f i ∉ m.monsterCells := by + simp [firstMonster_eq_none, ofFn_cells, List.mem_ofFn] + +/-- Reflecting a path. -/ +def Path.reflect (p : Path N) : Path N where + cells := p.cells.map Cell.reflect + nonempty := mt List.map_eq_nil_iff.1 p.nonempty + head_first_row := by + rw [List.head_map] + exact p.head_first_row + last_last_row := by + rw [List.getLast_map] + exact p.last_last_row + valid_move_seq := by + refine List.chain'_map_of_chain' _ ?_ p.valid_move_seq + intro x y h + simp_rw [Adjacent, Nat.dist, Cell.reflect, Fin.rev] at h ⊢ + omega + +lemma Path.firstMonster_reflect (p : Path N) (m : MonsterData N) : + p.reflect.firstMonster m.reflect = (p.firstMonster m).map Cell.reflect := by + simp_rw [firstMonster, reflect, List.find?_map] + convert rfl + simp only [Function.comp_apply, decide_eq_decide, MonsterData.monsterCells] + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · rcases h with ⟨i, hi⟩ + refine ⟨i, ?_⟩ + simpa [MonsterData.reflect, Cell.reflect, Prod.ext_iff] using hi + · rcases h with ⟨i, hi⟩ + refine ⟨i, ?_⟩ + simpa [MonsterData.reflect, Cell.reflect, Prod.ext_iff] using hi + +/-! ### API definitions and lemmas about `Strategy` -/ + +lemma Strategy.play_comp_castLE (s : Strategy N) (m : MonsterData N) {k₁ k₂ : ℕ} (hk : k₁ ≤ k₂) : + s.play m k₂ ∘ Fin.castLE hk = s.play m k₁ := by + induction hk + case refl => rfl + case step k' hk' hki => + rw [← hki, ← Fin.castLE_comp_castLE hk' (Nat.le_succ k'), ← Function.comp_assoc] + convert rfl + exact Fin.snoc_comp_castSucc.symm + +lemma Strategy.play_apply_of_le (s : Strategy N) (m : MonsterData N) {i k₁ k₂ : ℕ} (hi : i < k₁) + (hk : k₁ ≤ k₂) : s.play m k₂ ⟨i, hi.trans_le hk⟩ = s.play m k₁ ⟨i, hi⟩ := by + rw [← s.play_comp_castLE m hk] + rfl + +lemma Strategy.play_zero (s : Strategy N) (m : MonsterData N) {k : ℕ} (hk : 0 < k) : + s.play m k ⟨0, hk⟩ = (s Fin.elim0).firstMonster m := by + have hk' : 1 ≤ k := by omega + rw [s.play_apply_of_le m zero_lt_one hk'] + rfl + +lemma Strategy.play_one (s : Strategy N) (m : MonsterData N) {k : ℕ} (hk : 1 < k) : + s.play m k ⟨1, hk⟩ = (s ![(s Fin.elim0).firstMonster m]).firstMonster m := by + have hk' : 2 ≤ k := by omega + rw [s.play_apply_of_le m one_lt_two hk'] + simp only [play, Fin.snoc, lt_self_iff_false, ↓reduceDIte, Nat.reduceAdd, Nat.zero_eq, + Fin.mk_one, Fin.isValue, cast_eq, Nat.succ_eq_add_one] + congr + refine funext fun i ↦ ?_ + simp_rw [Fin.fin_one_eq_zero] + rfl + +lemma Strategy.play_two (s : Strategy N) (m : MonsterData N) {k : ℕ} (hk : 2 < k) : + s.play m k ⟨2, hk⟩ = (s ![(s Fin.elim0).firstMonster m, + (s ![(s Fin.elim0).firstMonster m]).firstMonster m]).firstMonster m := by + have hk' : 3 ≤ k := by omega + rw [s.play_apply_of_le m (by norm_num : 2 < 3) hk'] + simp only [play, Fin.snoc, lt_self_iff_false, ↓reduceDIte, Nat.reduceAdd, Nat.zero_eq, + Fin.mk_one, Fin.isValue, cast_eq, Nat.succ_eq_add_one] + congr + refine funext fun i ↦ ?_ + fin_cases i + · rfl + · have h : (1 : Fin 2) = Fin.last 1 := rfl + simp only [Fin.snoc_zero, Nat.reduceAdd, Fin.mk_one, Fin.isValue, Matrix.cons_val_one, + Matrix.head_cons] + simp only [h, Fin.snoc_last] + convert rfl + simp_rw [Fin.fin_one_eq_zero] + rfl + +lemma Strategy.WinsIn.mono (s : Strategy N) (m : MonsterData N) {k₁ k₂ : ℕ} (h : s.WinsIn m k₁) + (hk : k₁ ≤ k₂) : s.WinsIn m k₂ := by + refine Set.mem_of_mem_of_subset h ?_ + rw [Set.range_subset_range_iff_exists_comp] + exact ⟨Fin.castLE hk, (s.play_comp_castLE m hk).symm⟩ + +lemma Strategy.ForcesWinIn.mono (s : Strategy N) {k₁ k₂ : ℕ} (h : s.ForcesWinIn k₁) + (hk : k₁ ≤ k₂) : s.ForcesWinIn k₂ := + fun _ ↦ ((h _).mono) hk + +/-! ### Proof of lower bound with constructions used therein -/ + +/-- An arbitrary choice of monster positions, which is modified to put selected monsters in +desired places. -/ +def baseMonsterData (N : ℕ) : MonsterData N where + toFun := fun ⟨r, _, hrN⟩ ↦ ⟨↑r, by + rw [Fin.le_def] at hrN + rw [Nat.lt_add_one_iff] + exact hrN⟩ + inj' := fun ⟨⟨x, hx⟩, hx1, hxN⟩ ⟨⟨y, hy⟩, hy1, hyN⟩ h ↦ by + simp only [Fin.mk.injEq] at h + simpa using h + +/-- Positions for monsters with specified columns in the second and third rows (rows 1 and 2). -/ +def monsterData12 (hN : 2 ≤ N) (c₁ c₂ : Fin (N + 1)) : MonsterData N := + ((baseMonsterData N).setValue (row2 hN) c₂).setValue (row1 hN) c₁ + +lemma monsterData12_apply_row2 (hN : 2 ≤ N) {c₁ c₂ : Fin (N + 1)} (h : c₁ ≠ c₂) : + monsterData12 hN c₁ c₂ (row2 hN) = c₂ := by + rw [monsterData12, Function.Embedding.setValue_eq_of_ne] + · exact Function.Embedding.setValue_eq _ _ _ + · simp only [row1, row2, ne_eq, Subtype.mk.injEq] + simp only [Fin.ext_iff, Fin.val_one] + omega + · rw [Function.Embedding.setValue_eq] + exact h.symm + +lemma row1_mem_monsterCells_monsterData12 (hN : 2 ≤ N) (c₁ c₂ : Fin (N + 1)) : + (1, c₁) ∈ (monsterData12 hN c₁ c₂).monsterCells := by + exact Set.mem_range_self (row1 hN) + +lemma row2_mem_monsterCells_monsterData12 (hN : 2 ≤ N) {c₁ c₂ : Fin (N + 1)} (h : c₁ ≠ c₂) : + (⟨2, by omega⟩, c₂) ∈ (monsterData12 hN c₁ c₂).monsterCells := by + convert Set.mem_range_self (row2 hN) + exact (monsterData12_apply_row2 hN h).symm + +lemma Strategy.not_forcesWinIn_two (s : Strategy N) (hN : 2 ≤ N) : ¬ s.ForcesWinIn 2 := by + simp only [ForcesWinIn, WinsIn, Set.mem_range, not_forall, not_exists, Option.ne_none_iff_isSome] + let m1 : Cell N := (s Fin.elim0).findFstEq 1 + let m2 : Cell N := (s ![m1]).findFstEq 2 + let m : MonsterData N := monsterData12 hN m1.2 m2.2 + have h1r : m1.1 = 1 := Path.findFstEq_fst _ _ + have h2r : m2.1 = 2 := Path.findFstEq_fst _ _ + have h1 : m1 ∈ m.monsterCells := by + convert row1_mem_monsterCells_monsterData12 hN m1.2 m2.2 + have h2 : ((2 : Fin (N + 2)) : ℕ) = 2 := Nat.mod_eq_of_lt (by omega : 2 < N + 2) + refine ⟨m, fun i ↦ ?_⟩ + fin_cases i + · simp only [Strategy.play_zero, Path.firstMonster_eq_of_findFstEq_mem h1, Option.isSome_some] + · simp only [Strategy.play_one] + suffices ((s ![some m1]).firstMonster m).isSome = true by + rwa [Path.firstMonster_eq_of_findFstEq_mem h1] + simp_rw [m] + by_cases h : m1.2 = m2.2 + · rw [Path.firstMonster_isSome] + refine ⟨m1, ?_, h1⟩ + have h' : m1 = (⟨(((2 : Fin (N + 2)) : ℕ) - 1 : ℕ), by omega⟩, m2.2) := by + simp [h2, Prod.ext_iff, h1r, h2r, h] + nth_rw 2 [h'] + refine Path.findFstEq_fst_sub_one_mem _ ?_ + rw [ne_eq, Fin.ext_iff, h2] + norm_num + · rw [Path.firstMonster_isSome] + refine ⟨m2, Path.findFstEq_mem_cells _ _, ?_⟩ + convert row2_mem_monsterCells_monsterData12 hN h using 1 + simp [Prod.ext_iff, h2r, Fin.ext_iff, h2] + +lemma Strategy.ForcesWinIn.three_le {s : Strategy N} {k : ℕ} (hf : s.ForcesWinIn k) + (hN : 2 ≤ N) : 3 ≤ k := by + by_contra hk + exact s.not_forcesWinIn_two hN (Strategy.ForcesWinIn.mono s hf (by omega)) + +/-! ### Proof of upper bound and constructions used therein -/ + +/-- The first attempt in a winning strategy, as a function: first pass along the second row to +locate the monster there. -/ +def fn0 (N : ℕ) : Fin (2 * N + 2) → Cell N := + fun i ↦ if i = 0 then (0, 0) else + if h : (i : ℕ) < N + 2 then (1, ⟨(i : ℕ) - 1, by omega⟩) else + (⟨i - N, by omega⟩, ⟨N, by omega⟩) + +lemma injective_fn0 (N : ℕ) : Function.Injective (fn0 N) := by + intro a₁ a₂ + simp_rw [fn0] + split_ifs <;> simp [Prod.ext_iff, Fin.ext_iff] at * <;> omega + +/-- The first attempt in a winning strategy, as a `Path`. -/ +def path0 (hN : 2 ≤ N) : Path N := Path.ofFn (fn0 N) (by omega) (by simp [fn0]) + (by simp only [fn0, Fin.ext_iff]; split_ifs with h <;> simp at h ⊢ <;> omega) + (by + simp only [fn0] + intro i hi + split_ifs <;> simp [Adjacent, Nat.dist, Fin.ext_iff] at * <;> omega) + +/-- The second attempt in a winning strategy, as a function, if the monster in the second row +is not at an edge: pass to the left of that monster then along its column. -/ +def fn1OfNotEdge {c₁ : Fin (N + 1)} (hc₁ : c₁ ≠ 0) : Fin (N + 3) → Cell N := + fun i ↦ if h : (i : ℕ) ≤ 2 then (⟨(i : ℕ), by omega⟩, ⟨(c₁ : ℕ) - 1, by omega⟩) else + (⟨(i : ℕ) - 1, by omega⟩, c₁) + +/-- The second attempt in a winning strategy, as a function, if the monster in the second row +is not at an edge. -/ +def path1OfNotEdge {c₁ : Fin (N + 1)} (hc₁ : c₁ ≠ 0) : Path N := Path.ofFn (fn1OfNotEdge hc₁) + (by omega) (by simp [fn1OfNotEdge]) + (by simp only [fn1OfNotEdge, Fin.ext_iff]; split_ifs <;> simp; omega) + (by + simp only [fn1OfNotEdge] + intro i hi + rcases c₁ with ⟨c₁, hc₁N⟩ + rw [← Fin.val_ne_iff] at hc₁ + split_ifs <;> simp [Adjacent, Nat.dist, Fin.ext_iff] at * <;> omega) + +/-- The third attempt in a winning strategy, as a function, if the monster in the second row +is not at an edge: pass to the right of that monster then along its column. -/ +def fn2OfNotEdge {c₁ : Fin (N + 1)} (hc₁ : (c₁ : ℕ) ≠ N) : Fin (N + 3) → Cell N := + fun i ↦ if h : (i : ℕ) ≤ 2 then (⟨(i : ℕ), by omega⟩, ⟨(c₁ : ℕ) + 1, by omega⟩) else + (⟨(i : ℕ) - 1, by omega⟩, c₁) + +/-- The third attempt in a winning strategy, as a function, if the monster in the second row +is not at an edge. -/ +def path2OfNotEdge {c₁ : Fin (N + 1)} (hc₁ : (c₁ : ℕ) ≠ N) : Path N := Path.ofFn (fn2OfNotEdge hc₁) + (by omega) (by simp [fn2OfNotEdge]) + (by simp only [fn2OfNotEdge, Fin.ext_iff]; split_ifs <;> simp; omega) + (by + simp only [fn2OfNotEdge] + intro i hi + split_ifs <;> simp [Adjacent, Nat.dist, Fin.ext_iff] at * <;> omega) + +/-- The second attempt in a winning strategy, as a function, if the monster in the second row +is at the left edge: zigzag across the board so that, if we encounter a monster, we have a third +path we know will not encounter a monster. -/ +def fn1OfEdge0 (N : ℕ) : Fin (2 * N + 1) → Cell N := + fun i ↦ if h : (i : ℕ) = 2 * N then (⟨N + 1, by omega⟩, ⟨N, by omega⟩) else + (⟨((i : ℕ) + 1) / 2, by omega⟩, ⟨(i : ℕ) / 2 + 1, by omega⟩) + +/-- The second attempt in a winning strategy, as a `Path`, if the monster in the second row +is at the left edge. -/ +def path1OfEdge0 (hN : 2 ≤ N) : Path N := Path.ofFn (fn1OfEdge0 N) (by omega) + (by simp only [fn1OfEdge0, Fin.ext_iff]; split_ifs <;> simp; omega) + (by simp [fn1OfEdge0]) + (by + simp only [fn1OfEdge0] + intro i hi + split_ifs <;> simp [Adjacent, Nat.dist, Fin.ext_iff] at * <;> omega) + +/-- The second attempt in a winning strategy, as a `Path`, if the monster in the second row +is at the right edge. -/ +def path1OfEdgeN (hN : 2 ≤ N) : Path N := (path1OfEdge0 hN).reflect + +/-- The third attempt in a winning strategy, as a function, if the monster in the second row +is at the left edge and the second (zigzag) attempt encountered a monster: on the row before +the monster was encountered, move to the following row one place earlier, proceed to the left +edge and then along that column. -/ +def fn2OfEdge0 {r : Fin (N + 2)} (hr : (r : ℕ) ≤ N) : Fin (N + 2 * r - 1) → Cell N := + fun i ↦ if h₁ : (i : ℕ) + 2 < 2 * (r : ℕ) then + (⟨((i : ℕ) + 1) / 2, by omega⟩, ⟨(i : ℕ) / 2 + 1, by omega⟩) else + if h₂ : (i : ℕ) + 2 < 3 * (r : ℕ) then (r, ⟨3 * (r : ℕ) - 3 - (i : ℕ), by omega⟩) else + (⟨i + 3 - 2 * r, by omega⟩, 0) + +lemma fn2OfEdge0_apply_eq_fn1OfEdge0_apply_of_lt {r : Fin (N + 2)} (hr : (r : ℕ) ≤ N) {i : ℕ} + (h : i + 2 < 2 * (r : ℕ)) : fn2OfEdge0 hr ⟨i, by omega⟩ = fn1OfEdge0 N ⟨i, by omega⟩ := by + rw [fn1OfEdge0, fn2OfEdge0] + split_ifs with h₁ h₂ <;> simp [Fin.ext_iff, Prod.ext_iff] at * <;> omega + +/-- The third attempt in a winning strategy, as a `Path`, if the monster in the second row +is at the left edge and the second (zigzag) attempt encountered a monster. -/ +def path2OfEdge0 (hN : 2 ≤ N) {r : Fin (N + 2)} (hr2 : 2 ≤ (r : ℕ)) (hrN : (r : ℕ) ≤ N) : Path N := + Path.ofFn (fn2OfEdge0 hrN) (by omega) + (by simp only [fn2OfEdge0, Fin.ext_iff]; split_ifs <;> simp <;> omega) + (by simp only [fn2OfEdge0, Fin.ext_iff]; split_ifs <;> simp <;> omega) + (by + simp only [fn2OfEdge0] + intro i hi + split_ifs <;> simp [Adjacent, Nat.dist, Fin.ext_iff] at * <;> omega) + +/-- The third attempt in a winning strategy, as a `Path`, if the monster in the second row +is at the left edge and the second (zigzag) attempt encountered a monster, version that works +with junk values of `r` for convenience in defining the strategy before needing to prove things +about exactly where it can encounter monsters. -/ +def path2OfEdge0Def (hN : 2 ≤ N) (r : Fin (N + 2)) : Path N := + if h : 2 ≤ (r : ℕ) ∧ (r : ℕ) ≤ N then path2OfEdge0 hN h.1 h.2 else + path2OfEdge0 hN (r := ⟨2, by omega⟩) (le_refl _) hN + +/-- The third attempt in a winning strategy, as a `Path`, if the monster in the second row +is at the right edge and the second (zigzag) attempt encountered a monster. -/ +def path2OfEdgeNDef (hN : 2 ≤ N) (r : Fin (N + 2)) : Path N := + (path2OfEdge0Def hN r).reflect + +/-- The second attempt in a winning strategy, given the column of the monster in the second row, +as a `Path`. -/ +def path1 (hN : 2 ≤ N) (c₁ : Fin (N + 1)) : Path N := + if hc₁ : c₁ = 0 then path1OfEdge0 hN else + if (c₁ : ℕ) = N then path1OfEdgeN hN else + path1OfNotEdge hc₁ + +/-- The third attempt in a winning strategy, given the column of the monster in the second row +and the row of the monster (if any) found in the second attempt, as a `Path`. -/ +def path2 (hN : 2 ≤ N) (c₁ : Fin (N + 1)) (r : Fin (N + 2)) : Path N := + if c₁ = 0 then path2OfEdge0Def hN r else + if hc₁N : (c₁ : ℕ) = N then path2OfEdgeNDef hN r else + path2OfNotEdge hc₁N + +/-- A strategy that wins in three attempts. -/ +def winningStrategy (hN : 2 ≤ N) : Strategy N + | 0 => fun _ ↦ path0 hN + | 1 => fun r => path1 hN ((r 0).getD 0).2 + | _ + 2 => fun r => path2 hN ((r 0).getD 0).2 ((r 1).getD 0).1 + +lemma path0_firstMonster_eq_apply_row1 (hN : 2 ≤ N) (m : MonsterData N) : + (path0 hN).firstMonster m = some (1, m (row1 hN)) := by + simp_rw [path0, Path.firstMonster, Path.ofFn] + have h : (1, m (row1 hN)) = fn0 N ⟨(m (row1 hN) : ℕ) + 1, by omega⟩ := by + simp_rw [fn0] + split_ifs <;> simp [Prod.ext_iff, Fin.ext_iff] at *; omega + rw [h, List.find?_ofFn_eq_some_of_injective (injective_fn0 N)] + refine ⟨?_, fun j hj ↦ ?_⟩ + · rw [fn0] + split_ifs + · simp [Prod.ext_iff, Fin.ext_iff] at * + · have hm1 : (1, m (row1 hN)) ∈ m.monsterCells := Set.mem_range_self (row1 hN) + simpa [Prod.ext_iff, Fin.ext_iff] using hm1 + · rw [fn0] + split_ifs with h₁ + · simp [h₁, MonsterData.monsterCells] + · have hjN2 : (j : ℕ) < N + 2 := by + rw [Fin.lt_def] at hj + refine hj.trans ?_ + simp + simp only [MonsterData.monsterCells, h₁, ↓reduceIte, hjN2, ↓reduceDIte, Set.mem_range, + Prod.mk.injEq, Fin.ext_iff, Fin.val_one, Subtype.exists, Set.mem_Icc, exists_and_left, + decide_eq_true_eq, not_exists, not_and] + rintro i hi hix hij + simp only [Fin.ext_iff, Fin.val_zero] at h₁ + rw [eq_comm, Nat.sub_eq_iff_eq_add (by omega)] at hij + have hr : row1 hN = ⟨↑i, hix⟩ := by + simp_rw [Subtype.ext_iff, Fin.ext_iff, hi] + rfl + rw [Fin.lt_def, hij, add_lt_add_iff_right, Fin.val_fin_lt, hr] at hj + exact Nat.lt_irrefl _ hj + +lemma winningStrategy_play_one (hN : 2 ≤ N) (m : MonsterData N) : + (winningStrategy hN).play m 3 ⟨1, by norm_num⟩ = + (path1 hN (m (row1 hN))).firstMonster m := by + simp_rw [Strategy.play_one, winningStrategy, path0_firstMonster_eq_apply_row1] + rfl + +lemma winningStrategy_play_two (hN : 2 ≤ N) (m : MonsterData N) : + (winningStrategy hN).play m 3 ⟨2, by norm_num⟩ = + (path2 hN (m (row1 hN)) + (((path1 hN (m (row1 hN))).firstMonster m).getD 0).1).firstMonster m := by + simp_rw [Strategy.play_two, winningStrategy, path0_firstMonster_eq_apply_row1] + rfl + +lemma path1_firstMonster_of_not_edge (hN : 2 ≤ N) {m : MonsterData N} (hc₁0 : m (row1 hN) ≠ 0) + (hc₁N : (m (row1 hN) : ℕ) ≠ N) : + (path1 hN (m (row1 hN))).firstMonster m = none ∨ + (path1 hN (m (row1 hN))).firstMonster m = + some (⟨2, by omega⟩, ⟨(m (row1 hN) : ℕ) - 1, by omega⟩) := by + suffices h : ∀ c ∈ (path1 hN (m (row1 hN))).cells, c ∉ m.monsterCells ∨ + c = (⟨2, by omega⟩, ⟨(m (row1 hN) : ℕ) - 1, by omega⟩) by + simp only [Path.firstMonster] + by_cases hn : List.find? (fun x ↦ decide (x ∈ m.monsterCells)) + (path1 hN (m (row1 hN))).cells = none + · exact .inl hn + · rw [← ne_eq, Option.ne_none_iff_exists'] at hn + rcases hn with ⟨x, hx⟩ + have hxm := List.find?_some hx + have hx' := h _ (List.mem_of_find?_eq_some hx) + simp only [decide_eq_true_eq] at hxm + simp only [hxm, not_true_eq_false, false_or] at hx' + rw [hx'] at hx + exact .inr hx + simp only [path1, hc₁0, ↓reduceDIte, hc₁N, ↓reduceIte, path1OfNotEdge, Path.ofFn_cells, + List.forall_mem_ofFn_iff, fn1OfNotEdge] + intro j + split_ifs with h + · rcases j with ⟨j, hj⟩ + simp only at h + interval_cases j + · exact .inl (m.not_mem_monsterCells_of_fst_eq_zero rfl) + · simp only [Fin.mk_one, Prod.mk.injEq, Fin.ext_iff, Fin.val_one, OfNat.one_ne_ofNat, + and_true, or_false] + have hN' : 1 ≤ N := by omega + rw [m.mk_mem_monsterCells_iff_of_le (le_refl _) hN', ← ne_eq, ← Fin.val_ne_iff] + rw [← Fin.val_ne_iff] at hc₁0 + exact (Nat.sub_one_lt hc₁0).ne' + · exact .inr rfl + · refine .inl ?_ + simp only [MonsterData.monsterCells, Set.mem_range, Prod.mk.injEq, Fin.ext_iff, + EmbeddingLike.apply_eq_iff_eq, exists_eq_right, coe_coe_row1] + omega + +lemma path2_firstMonster_of_not_edge (hN : 2 ≤ N) {m : MonsterData N} (hc₁0 : m (row1 hN) ≠ 0) + (hc₁N : (m (row1 hN) : ℕ) ≠ N) (r : Fin (N + 2)) : + (path2 hN (m (row1 hN)) r).firstMonster m = none ∨ + (path2 hN (m (row1 hN)) r).firstMonster m = + some (⟨2, by omega⟩, ⟨(m (row1 hN) : ℕ) + 1, by omega⟩) := by + suffices h : ∀ c ∈ (path2 hN (m (row1 hN)) r).cells, c ∉ m.monsterCells ∨ + c = (⟨2, by omega⟩, ⟨(m (row1 hN) : ℕ) + 1, by omega⟩) by + simp only [Path.firstMonster] + by_cases hn : List.find? (fun x ↦ decide (x ∈ m.monsterCells)) + (path2 hN (m (row1 hN)) r).cells = none + · exact .inl hn + · rw [← ne_eq, Option.ne_none_iff_exists'] at hn + rcases hn with ⟨x, hx⟩ + have hxm := List.find?_some hx + have hx' := h _ (List.mem_of_find?_eq_some hx) + simp only [decide_eq_true_eq] at hxm + simp only [hxm, not_true_eq_false, false_or] at hx' + rw [hx'] at hx + exact .inr hx + simp only [path2, hc₁0, ↓reduceDIte, hc₁N, ↓reduceIte, path2OfNotEdge, Path.ofFn_cells, + List.forall_mem_ofFn_iff, fn2OfNotEdge] + intro j + split_ifs with h + · rcases j with ⟨j, hj⟩ + simp only at h + interval_cases j + · exact .inl (m.not_mem_monsterCells_of_fst_eq_zero rfl) + · simp only [Fin.mk_one, Prod.mk.injEq, Fin.ext_iff, Fin.val_one, OfNat.one_ne_ofNat, + and_true, or_false] + have hN' : 1 ≤ N := by omega + rw [m.mk_mem_monsterCells_iff_of_le (le_refl _) hN', ← ne_eq, ← Fin.val_ne_iff] + exact Nat.ne_add_one _ + · exact .inr rfl + · refine .inl ?_ + simp only [MonsterData.monsterCells, Set.mem_range, Prod.mk.injEq, Fin.ext_iff, + EmbeddingLike.apply_eq_iff_eq, exists_eq_right, coe_coe_row1] + omega + +lemma winningStrategy_play_one_eq_none_or_play_two_eq_none_of_not_edge (hN : 2 ≤ N) + {m : MonsterData N} (hc₁0 : m (row1 hN) ≠ 0) (hc₁N : (m (row1 hN) : ℕ) ≠ N) : + (winningStrategy hN).play m 3 ⟨1, by norm_num⟩ = none ∨ + (winningStrategy hN).play m 3 ⟨2, by norm_num⟩ = none := by + rw [winningStrategy_play_one, winningStrategy_play_two] + have h1 := path1_firstMonster_of_not_edge hN hc₁0 hc₁N + rcases h1 with h1 | h1 + · exact .inl h1 + · refine .inr ?_ + have h2 := path2_firstMonster_of_not_edge hN hc₁0 hc₁N + rcases h2 (((path1 hN (m (row1 hN))).firstMonster m).getD 0).1 with h2r | h2r + · exact h2r + · have h1' := (Path.mem_of_firstMonster_eq_some h1).2 + have h2' := (Path.mem_of_firstMonster_eq_some h2r).2 + have h12 : 1 ≤ (⟨2, by omega⟩ : Fin (N + 2)) := by + rw [Fin.le_def] + norm_num + rw [m.mk_mem_monsterCells_iff_of_le h12 hN] at h1' h2' + rw [h1', Fin.ext_iff] at h2' + simp at h2' + +lemma path2OfEdge0_firstMonster_eq_none_of_path1OfEdge0_firstMonster_eq_some (hN : 2 ≤ N) + {x : Cell N} (hx2 : 2 ≤ (x.1 : ℕ)) (hxN : (x.1 : ℕ) ≤ N) {m : MonsterData N} + (hc₁0 : m (row1 hN) = 0) (hx : (path1OfEdge0 hN).firstMonster m = some x) : + (path2OfEdge0 hN hx2 hxN).firstMonster m = none := by + rw [path2OfEdge0, Path.ofFn_firstMonster_eq_none] + rw [path1OfEdge0, Path.firstMonster, Path.ofFn_cells, List.find?_ofFn_eq_some] at hx + simp only [decide_eq_true_eq] at hx + rcases hx with ⟨hx, i, hix, hnm⟩ + have hi : (x.1 : ℕ) = ((i : ℕ) + 1) / 2 := by + rw [fn1OfEdge0] at hix + split_ifs at hix <;> simp [Prod.ext_iff, Fin.ext_iff] at hix <;> omega + have hi' : (i : ℕ) ≠ 2 * N := by + intro h + rw [← hix, fn1OfEdge0, dif_pos h] at hx + have h' := MonsterData.le_N_of_mem_monsterCells hx + simp at h' + intro j + by_cases h : (j : ℕ) + 2 < 2 * (x.1 : ℕ) + · rw [fn2OfEdge0_apply_eq_fn1OfEdge0_apply_of_lt hxN h] + simp_rw [Fin.lt_def] at hnm + refine hnm _ ?_ + simp only + omega + · rw [fn2OfEdge0, dif_neg h] + split_ifs with h' + · have hx1 : 1 ≤ x.1 := by + rw [Fin.le_def, Fin.val_one] + omega + rw [MonsterData.mk_mem_monsterCells_iff_of_le hx1 hxN] + rw [MonsterData.mem_monsterCells_iff_of_le hx1 hxN] at hx + simp_rw [hx, ← hix, fn1OfEdge0] + split_ifs <;> simp <;> omega + · rw [MonsterData.mk_mem_monsterCells_iff] + simp only [not_exists] + intro h1 hN' + rw [← hc₁0] + refine m.injective.ne ?_ + rw [← Subtype.coe_ne_coe, ← Fin.val_ne_iff, coe_coe_row1] + simp only + omega + +lemma winningStrategy_play_one_eq_none_or_play_two_eq_none_of_edge_zero (hN : 2 ≤ N) + {m : MonsterData N} (hc₁0 : m (row1 hN) = 0) : + (winningStrategy hN).play m 3 ⟨1, by norm_num⟩ = none ∨ + (winningStrategy hN).play m 3 ⟨2, by norm_num⟩ = none := by + rw [or_iff_not_imp_left] + intro h + rw [← ne_eq, Option.ne_none_iff_exists] at h + rcases h with ⟨x, hx⟩ + rw [winningStrategy_play_one] at hx + rw [winningStrategy_play_two, ← hx, Option.getD_some] + rw [path1, dif_pos hc₁0] at hx + have h1 := Path.mem_of_firstMonster_eq_some (hx.symm) + have hx2N : 2 ≤ (x.1 : ℕ) ∧ (x.1 : ℕ) ≤ N := by + rw [path1OfEdge0, Path.ofFn_cells, List.mem_ofFn] at h1 + rcases h1 with ⟨⟨i, rfl⟩, hm⟩ + refine ⟨?_, m.le_N_of_mem_monsterCells hm⟩ + rcases i with ⟨i, hi⟩ + have hi : 3 ≤ i := by + by_contra hi + interval_cases i <;> rw [fn1OfEdge0] at hm <;> split_ifs at hm with h + · simp at h + omega + · simp at hm + exact m.not_mem_monsterCells_of_fst_eq_zero rfl hm + · simp at h + omega + · dsimp only [Nat.reduceAdd, Nat.reduceDiv, Fin.mk_one] at hm + have h1N : 1 ≤ N := by omega + rw [m.mk_mem_monsterCells_iff_of_le (le_refl _) h1N] at hm + rw [row1, hm, Fin.ext_iff] at hc₁0 + simp at hc₁0 + · simp at h + omega + · simp only at hm + norm_num at hm + have h1N : 1 ≤ N := by omega + rw [m.mk_mem_monsterCells_iff_of_le (le_refl _) h1N] at hm + rw [row1, hm, Fin.ext_iff] at hc₁0 + simp at hc₁0 + rw [fn1OfEdge0] + split_ifs <;> simp <;> omega + rw [path2, if_pos hc₁0, path2OfEdge0Def, dif_pos hx2N] + exact path2OfEdge0_firstMonster_eq_none_of_path1OfEdge0_firstMonster_eq_some hN hx2N.1 + hx2N.2 hc₁0 hx.symm + +lemma winningStrategy_play_one_of_edge_N (hN : 2 ≤ N) {m : MonsterData N} + (hc₁N : (m (row1 hN) : ℕ) = N) : (winningStrategy hN).play m 3 ⟨1, by norm_num⟩ = + ((winningStrategy hN).play m.reflect 3 ⟨1, by norm_num⟩).map Cell.reflect := by + have hc₁0 : m (row1 hN) ≠ 0 := by + rw [← Fin.val_ne_iff, hc₁N, Fin.val_zero] + omega + have hc₁r0 : m.reflect (row1 hN) = 0 := by + simp only [MonsterData.reflect, Function.Embedding.coeFn_mk, Function.comp_apply, + ← Fin.rev_last, Fin.rev_inj] + rw [Fin.ext_iff] + exact hc₁N + simp_rw [winningStrategy_play_one hN, path1, path1OfEdgeN, dif_neg hc₁0, if_pos hc₁N, + dif_pos hc₁r0, ← Path.firstMonster_reflect, MonsterData.reflect_reflect] + +lemma winningStrategy_play_two_of_edge_N (hN : 2 ≤ N) {m : MonsterData N} + (hc₁N : (m (row1 hN) : ℕ) = N) : (winningStrategy hN).play m 3 ⟨2, by norm_num⟩ = + ((winningStrategy hN).play m.reflect 3 ⟨2, by norm_num⟩).map Cell.reflect := by + have hc₁0 : m (row1 hN) ≠ 0 := by + rw [← Fin.val_ne_iff, hc₁N, Fin.val_zero] + omega + have hc₁r0 : m.reflect (row1 hN) = 0 := by + simp only [MonsterData.reflect, Function.Embedding.coeFn_mk, Function.comp_apply, + ← Fin.rev_last, Fin.rev_inj] + rw [Fin.ext_iff] + exact hc₁N + simp_rw [winningStrategy_play_two hN, path1, path1OfEdgeN, path2, path2OfEdgeNDef, if_neg hc₁0, + dif_neg hc₁0, if_pos hc₁N, dif_pos hc₁N, if_pos hc₁r0, dif_pos hc₁r0, + ← Path.firstMonster_reflect, MonsterData.reflect_reflect] + convert rfl using 4 + nth_rw 2 [← m.reflect_reflect] + rw [Path.firstMonster_reflect] + rcases ((path1OfEdge0 hN).firstMonster m.reflect).eq_none_or_eq_some with h | h + · simp [h] + · rcases h with ⟨x, hx⟩ + simp [hx, Cell.reflect] + +lemma winningStrategy_play_one_eq_none_or_play_two_eq_none_of_edge_N (hN : 2 ≤ N) + {m : MonsterData N} (hc₁N : (m (row1 hN) : ℕ) = N) : + (winningStrategy hN).play m 3 ⟨1, by norm_num⟩ = none ∨ + (winningStrategy hN).play m 3 ⟨2, by norm_num⟩ = none := by + simp_rw [winningStrategy_play_one_of_edge_N hN hc₁N, winningStrategy_play_two_of_edge_N hN hc₁N, + Option.map_eq_none'] + have hc₁r0 : m.reflect (row1 hN) = 0 := by + simp only [MonsterData.reflect, Function.Embedding.coeFn_mk, Function.comp_apply, + ← Fin.rev_last, Fin.rev_inj] + rw [Fin.ext_iff] + exact hc₁N + exact winningStrategy_play_one_eq_none_or_play_two_eq_none_of_edge_zero hN hc₁r0 + +lemma winningStrategy_play_one_eq_none_or_play_two_eq_none (hN : 2 ≤ N) (m : MonsterData N) : + (winningStrategy hN).play m 3 ⟨1, by norm_num⟩ = none ∨ + (winningStrategy hN).play m 3 ⟨2, by norm_num⟩ = none := by + by_cases hc₁0 : m (row1 hN) = 0 + · exact winningStrategy_play_one_eq_none_or_play_two_eq_none_of_edge_zero hN hc₁0 + · by_cases hc₁N : (m (row1 hN) : ℕ) = N + · exact winningStrategy_play_one_eq_none_or_play_two_eq_none_of_edge_N hN hc₁N + · exact winningStrategy_play_one_eq_none_or_play_two_eq_none_of_not_edge hN hc₁0 hc₁N + +lemma winningStrategy_forcesWinIn_three (hN : 2 ≤ N) : + (winningStrategy hN).ForcesWinIn 3 := by + intro m + rcases winningStrategy_play_one_eq_none_or_play_two_eq_none hN m with h | h + · rw [Strategy.WinsIn] + convert Set.mem_range_self (⟨1, by norm_num⟩ : Fin 3) + exact h.symm + · rw [Strategy.WinsIn] + convert Set.mem_range_self (⟨2, by norm_num⟩ : Fin 3) + exact h.symm + +/-- This is to be determined by the solver of the original problem (and much of the difficulty +for contestants is in finding this answer rather than spending time attempting to prove a +conjectured answer around log₂ 2023). -/ +def answer : ℕ := 3 + +/-- The final result, combining upper and lower bounds. -/ +theorem result : IsLeast {k | ∃ s : Strategy 2022, s.ForcesWinIn k} answer := by + simp_rw [IsLeast, mem_lowerBounds, Set.mem_setOf, forall_exists_index] + exact ⟨⟨winningStrategy (by norm_num), winningStrategy_forcesWinIn_three (by norm_num)⟩, + fun k s h ↦ h.three_le (by norm_num)⟩ + +end Imo2024Q5 From 0b09ad8797613f7bb008e4223aca9537e3402185 Mon Sep 17 00:00:00 2001 From: Vincent Beffara Date: Sat, 16 Nov 2024 08:30:43 +0000 Subject: [PATCH 010/829] feat(Topology/ContinuousMap/Interval): define the concatenation operation on continuous maps on intervals as a continuous map (#17621) The main definition is a continuous map `transCM` on pairs of compatible continuous maps from the intervals `Icc a b` and `Icc b c` to form a continuous map on `Icc a c`. --- Mathlib.lean | 1 + Mathlib/Topology/ContinuousMap/Basic.lean | 5 + Mathlib/Topology/ContinuousMap/Interval.lean | 146 +++++++++++++++++++ 3 files changed, 152 insertions(+) create mode 100644 Mathlib/Topology/ContinuousMap/Interval.lean diff --git a/Mathlib.lean b/Mathlib.lean index 64355d7319159..0c70db84364d4 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4985,6 +4985,7 @@ import Mathlib.Topology.ContinuousMap.CompactlySupported import Mathlib.Topology.ContinuousMap.ContinuousMapZero import Mathlib.Topology.ContinuousMap.Defs import Mathlib.Topology.ContinuousMap.Ideals +import Mathlib.Topology.ContinuousMap.Interval import Mathlib.Topology.ContinuousMap.Lattice import Mathlib.Topology.ContinuousMap.LocallyConstant import Mathlib.Topology.ContinuousMap.Ordered diff --git a/Mathlib/Topology/ContinuousMap/Basic.lean b/Mathlib/Topology/ContinuousMap/Basic.lean index a09f0d851bd60..eb8f143e4e116 100644 --- a/Mathlib/Topology/ContinuousMap/Basic.lean +++ b/Mathlib/Topology/ContinuousMap/Basic.lean @@ -350,6 +350,11 @@ theorem liftCover_restrict' {s : Set α} {hs : s ∈ A} : end Gluing +/-- `Set.inclusion` as a bundled continuous map. -/ +def inclusion {s t : Set α} (h : s ⊆ t) : C(s, t) where + toFun := Set.inclusion h + continuous_toFun := continuous_inclusion h + end ContinuousMap section Lift diff --git a/Mathlib/Topology/ContinuousMap/Interval.lean b/Mathlib/Topology/ContinuousMap/Interval.lean new file mode 100644 index 0000000000000..b1053a5b0c33c --- /dev/null +++ b/Mathlib/Topology/ContinuousMap/Interval.lean @@ -0,0 +1,146 @@ +/- +Copyright (c) 2024 Vincent Beffara. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Vincent Beffara +-/ +import Mathlib.Topology.CompactOpen +import Mathlib.Topology.Order.ProjIcc + +/-! +# Continuous bundled maps on intervals + +In this file we prove a few results about `ContinuousMap` when the domain is an interval. +-/ + +open Set ContinuousMap Filter Topology + +namespace ContinuousMap + +variable {α : Type*} [LinearOrder α] [TopologicalSpace α] [OrderTopology α] +variable {a b c : α} [Fact (a ≤ b)] [Fact (b ≤ c)] +variable {E : Type*} [TopologicalSpace E] + +/-- The embedding into an interval from a sub-interval lying on the left, as a `ContinuousMap`. -/ +def IccInclusionLeft : C(Icc a b, Icc a c) := + .inclusion <| Icc_subset_Icc le_rfl Fact.out + +/-- The embedding into an interval from a sub-interval lying on the right, as a `ContinuousMap`. -/ +def IccInclusionRight : C(Icc b c, Icc a c) := + .inclusion <| Icc_subset_Icc Fact.out le_rfl + +/-- The map `projIcc` from `α` onto an interval in `α`, as a `ContinuousMap`. -/ +def projIccCM : C(α, Icc a b) := + ⟨projIcc a b Fact.out, continuous_projIcc⟩ + +/-- The extension operation from continuous maps on an interval to continuous maps on the whole + type, as a `ContinuousMap`. -/ +def IccExtendCM : C(C(Icc a b, E), C(α, E)) where + toFun f := f.comp projIccCM + continuous_toFun := continuous_precomp projIccCM + +@[simp] +theorem IccExtendCM_of_mem {f : C(Icc a b, E)} {x : α} (hx : x ∈ Icc a b) : + IccExtendCM f x = f ⟨x, hx⟩ := by + simp [IccExtendCM, projIccCM, projIcc, hx.1, hx.2] + +/-- The concatenation of two continuous maps defined on adjacent intervals. If the values of the +functions on the common bound do not agree, this is defined as an arbitrarily chosen constant +map. See `concatCM` for the corresponding map on the subtype of compatible function pairs. -/ +noncomputable def concat (f : C(Icc a b, E)) (g : C(Icc b c, E)) : + C(Icc a c, E) := by + by_cases hb : f ⊤ = g ⊥ + · let h (t : α) : E := if t ≤ b then IccExtendCM f t else IccExtendCM g t + suffices Continuous h from ⟨fun t => h t, by fun_prop⟩ + apply Continuous.if_le (by fun_prop) (by fun_prop) continuous_id continuous_const + rintro x rfl + simpa [IccExtendCM, projIccCM] + · exact .const _ (f ⊥) -- junk value + +variable {f : C(Icc a b, E)} {g : C(Icc b c, E)} + +theorem concat_comp_IccInclusionLeft (hb : f ⊤ = g ⊥) : + (concat f g).comp IccInclusionLeft = f := by + ext x + simp [concat, IccExtendCM, hb, IccInclusionLeft, projIccCM, inclusion, x.2.2] + +theorem concat_comp_IccInclusionRight (hb : f ⊤ = g ⊥) : + (concat f g).comp IccInclusionRight = g := by + ext ⟨x, hx⟩ + obtain rfl | hxb := eq_or_ne x b + · simpa [concat, IccInclusionRight, IccExtendCM, projIccCM, inclusion, hb] + · have h : ¬ x ≤ b := lt_of_le_of_ne hx.1 (Ne.symm hxb) |>.not_le + simp [concat, hb, IccInclusionRight, h, IccExtendCM, projIccCM, projIcc, inclusion, hx.2, hx.1] + +@[simp] +theorem concat_left (hb : f ⊤ = g ⊥) {t : Icc a c} (ht : t ≤ b) : + concat f g t = f ⟨t, t.2.1, ht⟩ := by + nth_rewrite 2 [← concat_comp_IccInclusionLeft hb] + rfl + +@[simp] +theorem concat_right (hb : f ⊤ = g ⊥) {t : Icc a c} (ht : b ≤ t) : + concat f g t = g ⟨t, ht, t.2.2⟩ := by + nth_rewrite 2 [← concat_comp_IccInclusionRight hb] + rfl + +theorem tendsto_concat {ι : Type*} {p : Filter ι} {F : ι → C(Icc a b, E)} {G : ι → C(Icc b c, E)} + (hfg : ∀ᶠ i in p, (F i) ⊤ = (G i) ⊥) (hfg' : f ⊤ = g ⊥) + (hf : Tendsto F p (𝓝 f)) (hg : Tendsto G p (𝓝 g)) : + Tendsto (fun i => concat (F i) (G i)) p (𝓝 (concat f g)) := by + rw [tendsto_nhds_compactOpen] at hf hg ⊢ + rintro K hK U hU hfgU + have h : b ∈ Icc a c := ⟨Fact.out, Fact.out⟩ + let K₁ : Set (Icc a b) := projIccCM '' (Subtype.val '' (K ∩ Iic ⟨b, h⟩)) + let K₂ : Set (Icc b c) := projIccCM '' (Subtype.val '' (K ∩ Ici ⟨b, h⟩)) + have hK₁ : IsCompact K₁ := + hK.inter_right isClosed_Iic |>.image continuous_subtype_val |>.image projIccCM.continuous + have hK₂ : IsCompact K₂ := + hK.inter_right isClosed_Ici |>.image continuous_subtype_val |>.image projIccCM.continuous + have hfU : MapsTo f K₁ U := by + rw [← concat_comp_IccInclusionLeft hfg'] + apply hfgU.comp + rintro x ⟨y, ⟨⟨z, hz⟩, ⟨h1, (h2 : z ≤ b)⟩, rfl⟩, rfl⟩ + simpa [projIccCM, projIcc, h2, hz.1] using h1 + have hgU : MapsTo g K₂ U := by + rw [← concat_comp_IccInclusionRight hfg'] + apply hfgU.comp + rintro x ⟨y, ⟨⟨z, hz⟩, ⟨h1, (h2 : b ≤ z)⟩, rfl⟩, rfl⟩ + simpa [projIccCM, projIcc, h2, hz.2] using h1 + filter_upwards [hf K₁ hK₁ U hU hfU, hg K₂ hK₂ U hU hgU, hfg] with i hf hg hfg x hx + by_cases hxb : x ≤ b + · rw [concat_left hfg hxb] + refine hf ⟨x, ⟨x, ⟨hx, hxb⟩, rfl⟩, ?_⟩ + simp [projIccCM, projIcc, hxb, x.2.1] + · replace hxb : b ≤ x := lt_of_not_le hxb |>.le + rw [concat_right hfg hxb] + refine hg ⟨x, ⟨x, ⟨hx, hxb⟩, rfl⟩, ?_⟩ + simp [projIccCM, projIcc, hxb, x.2.2] + +/-- The concatenation of compatible pairs of continuous maps on adjacent intervals, defined as a +`ContinuousMap` on a subtype of the product. -/ +noncomputable def concatCM : + C({fg : C(Icc a b, E) × C(Icc b c, E) // fg.1 ⊤ = fg.2 ⊥}, C(Icc a c, E)) + where + toFun fg := concat fg.val.1 fg.val.2 + continuous_toFun := by + let S : Set (C(Icc a b, E) × C(Icc b c, E)) := {fg | fg.1 ⊤ = fg.2 ⊥} + change Continuous (S.restrict concat.uncurry) + refine continuousOn_iff_continuous_restrict.mp (fun fg hfg => ?_) + refine tendsto_concat ?_ hfg ?_ ?_ + · exact eventually_nhdsWithin_of_forall (fun _ => id) + · exact tendsto_nhdsWithin_of_tendsto_nhds continuousAt_fst + · exact tendsto_nhdsWithin_of_tendsto_nhds continuousAt_snd + +@[simp] +theorem concatCM_left {x : Icc a c} (hx : x ≤ b) + {fg : {fg : C(Icc a b, E) × C(Icc b c, E) // fg.1 ⊤ = fg.2 ⊥}} : + concatCM fg x = fg.1.1 ⟨x.1, x.2.1, hx⟩ := by + exact concat_left fg.2 hx + +@[simp] +theorem concatCM_right {x : Icc a c} (hx : b ≤ x) + {fg : {fg : C(Icc a b, E) × C(Icc b c, E) // fg.1 ⊤ = fg.2 ⊥}} : + concatCM fg x = fg.1.2 ⟨x.1, hx, x.2.2⟩ := + concat_right fg.2 hx + +end ContinuousMap From 9a9cb1894400637ebac1c1c8ac2702c23412ee5e Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Sat, 16 Nov 2024 08:30:44 +0000 Subject: [PATCH 011/829] chore(Data/Set): split Finite.lean into `Defs`, `Basic`, `Lemmas` (#18619) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit `Data.Set.Finite` had a lot of imports and in turn was imported often. So let's split it up along import lines: * `Data/Finite/Defs.lean` now contains the definition of `Set.Finite` * `Data/Set/Finite/Basic.lean` knows only `Set.Finite` and `Finset` * `Data/Finite/Prod.lean` now contains finiteness results of products of sets * `Data/Fintype/Pi.lean` now contains finiteness results of sets of functions * `Data/Set/Finite/Lattice.lean` contains finiteness results of lattice operations on sets * `Data/Set/Finite/List.lean` contains finiteness results on lists converted to sets * `Data/Set/Finite/Monad.lean` also knows about the monad structure on `Set` * `Data/Set/Finite/Powerset.lean` contains finiteness results on powersets * `Data/Set/Finite/Range.lean` contains finiteness results on `Set.range` * `Data/Set/Finite/Lemmas.lean` is for lemmas *about* `Set.Finite` (and therefore does not contain any `Finite` instances itself). Co-authored-by: Anne Baanen Co-authored-by: Yaël Dillies --- Archive/Imo/Imo1969Q1.lean | 2 +- Archive/Imo/Imo2008Q2.lean | 4 +- Archive/Wiedijk100Theorems/CubingACube.lean | 2 +- Mathlib.lean | 8 +- Mathlib/Algebra/BigOperators/Finprod.lean | 1 + .../Algebra/BigOperators/Group/Finset.lean | 2 + Mathlib/Algebra/BigOperators/Ring.lean | 1 + Mathlib/Algebra/Group/ConjFinite.lean | 1 - Mathlib/Algebra/Group/FiniteSupport.lean | 2 +- Mathlib/Algebra/Group/Int.lean | 1 - .../Algebra/Group/Pointwise/Finset/Basic.lean | 1 + Mathlib/Algebra/Module/Submodule/Range.lean | 1 + Mathlib/Algebra/Order/Rearrangement.lean | 1 + Mathlib/Algebra/Polynomial/Roots.lean | 2 + Mathlib/Algebra/Ring/Subsemiring/Basic.lean | 1 - Mathlib/Algebra/Star/Pointwise.lean | 2 +- .../SplitSimplicialObject.lean | 1 + Mathlib/Analysis/Normed/Module/Dual.lean | 4 +- Mathlib/CategoryTheory/Types.lean | 2 +- Mathlib/Combinatorics/HalesJewett.lean | 1 + Mathlib/Combinatorics/Hall/Finite.lean | 3 +- Mathlib/Combinatorics/SetFamily/Shatter.lean | 1 + Mathlib/Combinatorics/SimpleGraph/Basic.lean | 3 +- Mathlib/Combinatorics/SimpleGraph/Clique.lean | 1 + Mathlib/Combinatorics/SimpleGraph/Dart.lean | 1 + Mathlib/Combinatorics/SimpleGraph/Finite.lean | 1 + Mathlib/Combinatorics/SimpleGraph/Maps.lean | 1 + .../Combinatorics/SimpleGraph/Matching.lean | 1 + .../SimpleGraph/StronglyRegular.lean | 1 - Mathlib/Data/DFinsupp/Defs.lean | 2 +- Mathlib/Data/Finite/Defs.lean | 96 ++- Mathlib/Data/Finite/Prod.lean | 172 +++++ Mathlib/Data/Finset/Grade.lean | 3 +- Mathlib/Data/Finset/Interval.lean | 1 + Mathlib/Data/Finset/NAry.lean | 4 +- Mathlib/Data/Finset/Preimage.lean | 3 +- Mathlib/Data/Finset/Slice.lean | 1 + Mathlib/Data/Finsupp/Defs.lean | 3 +- Mathlib/Data/Fintype/CardEmbedding.lean | 1 + Mathlib/Data/Fintype/Fin.lean | 1 + Mathlib/Data/Fintype/Order.lean | 4 +- Mathlib/Data/Fintype/Pi.lean | 44 ++ Mathlib/Data/MLList/BestFirst.lean | 2 +- Mathlib/Data/Matrix/Basic.lean | 1 + Mathlib/Data/Matroid/Basic.lean | 3 +- Mathlib/Data/Nat/PrimeFin.lean | 2 +- Mathlib/Data/Set/Countable.lean | 4 +- .../Set/{Finite.lean => Finite/Basic.lean} | 645 +----------------- Mathlib/Data/Set/Finite/Lattice.lean | 399 +++++++++++ Mathlib/Data/Set/Finite/Lemmas.lean | 92 +++ Mathlib/Data/Set/Finite/List.lean | 35 + Mathlib/Data/Set/Finite/Monad.lean | 125 ++++ Mathlib/Data/Set/Finite/Powerset.lean | 56 ++ Mathlib/Data/Set/Finite/Range.lean | 97 +++ Mathlib/Data/Set/MemPartition.lean | 2 +- Mathlib/Data/Set/Pointwise/Finite.lean | 2 +- Mathlib/Data/Setoid/Partition.lean | 2 +- Mathlib/GroupTheory/GroupAction/Basic.lean | 5 +- Mathlib/GroupTheory/GroupAction/Quotient.lean | 9 +- Mathlib/GroupTheory/Perm/ClosureSwap.lean | 1 - Mathlib/GroupTheory/Perm/Sign.lean | 1 + Mathlib/GroupTheory/QuotientGroup/Finite.lean | 1 + Mathlib/LinearAlgebra/Multilinear/Basic.lean | 1 + Mathlib/MeasureTheory/SetAlgebra.lean | 1 + Mathlib/ModelTheory/FinitelyGenerated.lean | 1 + Mathlib/Order/Atoms/Finite.lean | 1 + Mathlib/Order/Category/NonemptyFinLinOrd.lean | 9 +- Mathlib/Order/CompleteSublattice.lean | 1 + .../ConditionallyCompleteLattice/Finset.lean | 3 +- Mathlib/Order/Filter/AtTopBot.lean | 1 + Mathlib/Order/Filter/Basic.lean | 2 +- Mathlib/Order/Filter/Cofinite.lean | 2 + Mathlib/Order/Interval/Set/Infinite.lean | 3 +- Mathlib/Order/KonigLemma.lean | 1 + Mathlib/Order/PartialSups.lean | 4 +- Mathlib/Order/Partition/Equipartition.lean | 1 + Mathlib/Order/Partition/Finpartition.lean | 2 +- .../Order/SuccPred/LinearLocallyFinite.lean | 1 + Mathlib/Order/SupClosed.lean | 3 +- Mathlib/Order/UpperLower/LocallyFinite.lean | 1 + Mathlib/RingTheory/Finiteness/Defs.lean | 1 + Mathlib/RingTheory/Localization/Defs.lean | 1 + Mathlib/RingTheory/MvPolynomial/Basic.lean | 1 + .../NonUnitalSubsemiring/Basic.lean | 1 + Mathlib/RingTheory/Polynomial/Dickson.lean | 2 +- Mathlib/SetTheory/Cardinal/Basic.lean | 1 + 86 files changed, 1228 insertions(+), 689 deletions(-) rename Mathlib/Data/Set/{Finite.lean => Finite/Basic.lean} (57%) create mode 100644 Mathlib/Data/Set/Finite/Lattice.lean create mode 100644 Mathlib/Data/Set/Finite/Lemmas.lean create mode 100644 Mathlib/Data/Set/Finite/List.lean create mode 100644 Mathlib/Data/Set/Finite/Monad.lean create mode 100644 Mathlib/Data/Set/Finite/Powerset.lean create mode 100644 Mathlib/Data/Set/Finite/Range.lean diff --git a/Archive/Imo/Imo1969Q1.lean b/Archive/Imo/Imo1969Q1.lean index c0d69cc499b23..9b8338fec35bc 100644 --- a/Archive/Imo/Imo1969Q1.lean +++ b/Archive/Imo/Imo1969Q1.lean @@ -5,8 +5,8 @@ Authors: Kevin Lacker -/ import Mathlib.Algebra.Ring.Identities import Mathlib.Data.Int.NatPrime +import Mathlib.Data.Set.Finite.Lemmas import Mathlib.Tactic.Linarith -import Mathlib.Data.Set.Finite /-! # IMO 1969 Q1 diff --git a/Archive/Imo/Imo2008Q2.lean b/Archive/Imo/Imo2008Q2.lean index ff333458edede..0340c3f70479b 100644 --- a/Archive/Imo/Imo2008Q2.lean +++ b/Archive/Imo/Imo2008Q2.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Manuel Candales -/ import Mathlib.Data.Real.Basic -import Mathlib.Data.Set.Finite -import Mathlib.Tactic.FieldSimp +import Mathlib.Data.Set.Finite.Lattice import Mathlib.Tactic.Abel +import Mathlib.Tactic.FieldSimp import Mathlib.Tactic.Linarith import Mathlib.Tactic.Ring diff --git a/Archive/Wiedijk100Theorems/CubingACube.lean b/Archive/Wiedijk100Theorems/CubingACube.lean index b0497c13dc2e4..6a9ccd178470c 100644 --- a/Archive/Wiedijk100Theorems/CubingACube.lean +++ b/Archive/Wiedijk100Theorems/CubingACube.lean @@ -5,7 +5,7 @@ Authors: Floris van Doorn -/ import Mathlib.Algebra.Order.Interval.Set.Group import Mathlib.Data.Real.Basic -import Mathlib.Data.Set.Finite +import Mathlib.Data.Set.Finite.Lemmas import Mathlib.Order.Interval.Set.Disjoint /-! diff --git a/Mathlib.lean b/Mathlib.lean index 0c70db84364d4..679585a362a5a 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2776,7 +2776,13 @@ import Mathlib.Data.Set.Countable import Mathlib.Data.Set.Defs import Mathlib.Data.Set.Enumerate import Mathlib.Data.Set.Equitable -import Mathlib.Data.Set.Finite +import Mathlib.Data.Set.Finite.Basic +import Mathlib.Data.Set.Finite.Lattice +import Mathlib.Data.Set.Finite.Lemmas +import Mathlib.Data.Set.Finite.List +import Mathlib.Data.Set.Finite.Monad +import Mathlib.Data.Set.Finite.Powerset +import Mathlib.Data.Set.Finite.Range import Mathlib.Data.Set.Function import Mathlib.Data.Set.Functor import Mathlib.Data.Set.Image diff --git a/Mathlib/Algebra/BigOperators/Finprod.lean b/Mathlib/Algebra/BigOperators/Finprod.lean index 1248df966b693..9205acc2d2800 100644 --- a/Mathlib/Algebra/BigOperators/Finprod.lean +++ b/Mathlib/Algebra/BigOperators/Finprod.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.Group.FiniteSupport import Mathlib.Algebra.NoZeroSMulDivisors.Basic import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Algebra.Order.Ring.Defs +import Mathlib.Data.Set.Finite.Lattice import Mathlib.Data.Set.Subsingleton /-! diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index 2c466c33831da..29817d8370d97 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -5,7 +5,9 @@ Authors: Johannes Hölzl -/ import Mathlib.Algebra.Group.Indicator import Mathlib.Data.Finset.Piecewise +import Mathlib.Data.Finset.Powerset import Mathlib.Data.Finset.Preimage +import Mathlib.Data.Fintype.Pi /-! # Big operators diff --git a/Mathlib/Algebra/BigOperators/Ring.lean b/Mathlib/Algebra/BigOperators/Ring.lean index a25c14300bd4b..ff6ed790f07e8 100644 --- a/Mathlib/Algebra/BigOperators/Ring.lean +++ b/Mathlib/Algebra/BigOperators/Ring.lean @@ -6,6 +6,7 @@ Authors: Johannes Hölzl import Mathlib.Algebra.BigOperators.GroupWithZero.Finset import Mathlib.Algebra.BigOperators.Ring.Multiset import Mathlib.Algebra.Field.Defs +import Mathlib.Data.Finset.Max import Mathlib.Data.Fintype.Powerset import Mathlib.Data.Int.Cast.Lemmas diff --git a/Mathlib/Algebra/Group/ConjFinite.lean b/Mathlib/Algebra/Group/ConjFinite.lean index aab17a8c4e97a..75b02165e48f9 100644 --- a/Mathlib/Algebra/Group/ConjFinite.lean +++ b/Mathlib/Algebra/Group/ConjFinite.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Rodriguez -/ import Mathlib.Algebra.Group.Conj -import Mathlib.Data.Finite.Prod import Mathlib.Data.Fintype.Units /-! diff --git a/Mathlib/Algebra/Group/FiniteSupport.lean b/Mathlib/Algebra/Group/FiniteSupport.lean index 7e4209101e421..d1e51c56165ce 100644 --- a/Mathlib/Algebra/Group/FiniteSupport.lean +++ b/Mathlib/Algebra/Group/FiniteSupport.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash -/ import Mathlib.Algebra.Group.Support -import Mathlib.Data.Set.Finite +import Mathlib.Data.Set.Finite.Basic /-! # Finiteness of support diff --git a/Mathlib/Algebra/Group/Int.lean b/Mathlib/Algebra/Group/Int.lean index 478a88a4799ad..1fcf3f10ef98c 100644 --- a/Mathlib/Algebra/Group/Int.lean +++ b/Mathlib/Algebra/Group/Int.lean @@ -17,7 +17,6 @@ See note [foundational algebra order theory]. assert_not_exists Ring assert_not_exists DenselyOrdered -assert_not_exists Set.range open Nat diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index 231d6b2463298..a7cd803e05c07 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -6,6 +6,7 @@ Authors: Floris van Doorn, Yaël Dillies import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Algebra.Group.Action.Pi import Mathlib.Data.Finset.Density +import Mathlib.Data.Finset.Max import Mathlib.Data.Finset.NAry import Mathlib.Data.Set.Pointwise.Finite import Mathlib.Data.Set.Pointwise.ListOfFn diff --git a/Mathlib/Algebra/Module/Submodule/Range.lean b/Mathlib/Algebra/Module/Submodule/Range.lean index b55c38d7fb24f..c2117e87b75a0 100644 --- a/Mathlib/Algebra/Module/Submodule/Range.lean +++ b/Mathlib/Algebra/Module/Submodule/Range.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl, Mario Carneiro, Kevin Buzzard, Yury Kudryashov, Fréd Heather Macbeth -/ import Mathlib.Algebra.Module.Submodule.Ker +import Mathlib.Data.Set.Finite.Range /-! # Range of linear maps diff --git a/Mathlib/Algebra/Order/Rearrangement.lean b/Mathlib/Algebra/Order/Rearrangement.lean index f69acca4d99d3..a6c4bfabc5fbd 100644 --- a/Mathlib/Algebra/Order/Rearrangement.lean +++ b/Mathlib/Algebra/Order/Rearrangement.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.Order.Module.OrderedSMul import Mathlib.Algebra.Order.Module.Synonym import Mathlib.Data.Prod.Lex import Mathlib.Data.Set.Image +import Mathlib.Data.Finset.Max import Mathlib.GroupTheory.Perm.Support import Mathlib.Order.Monotone.Monovary import Mathlib.Tactic.Abel diff --git a/Mathlib/Algebra/Polynomial/Roots.lean b/Mathlib/Algebra/Polynomial/Roots.lean index 3db1a24da52f9..ac9cd6e553c25 100644 --- a/Mathlib/Algebra/Polynomial/Roots.lean +++ b/Mathlib/Algebra/Polynomial/Roots.lean @@ -5,6 +5,8 @@ Authors: Chris Hughes, Johannes Hölzl, Kim Morrison, Jens Wagemaker, Johan Comm -/ import Mathlib.Algebra.Polynomial.BigOperators import Mathlib.Algebra.Polynomial.RingDivision +import Mathlib.Data.Fintype.Pi +import Mathlib.Data.Set.Finite.Lemmas import Mathlib.RingTheory.Localization.FractionRing import Mathlib.SetTheory.Cardinal.Basic diff --git a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean index 4a94131046f41..c98fc5fd857a1 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean @@ -9,7 +9,6 @@ import Mathlib.Algebra.Ring.Action.Subobjects import Mathlib.Algebra.Ring.Equiv import Mathlib.Algebra.Ring.Prod import Mathlib.Algebra.Ring.Subsemiring.Defs -import Mathlib.Data.Set.Finite import Mathlib.GroupTheory.Submonoid.Centralizer import Mathlib.RingTheory.NonUnitalSubsemiring.Basic diff --git a/Mathlib/Algebra/Star/Pointwise.lean b/Mathlib/Algebra/Star/Pointwise.lean index 6c37067eee6ab..75524536ac45c 100644 --- a/Mathlib/Algebra/Star/Pointwise.lean +++ b/Mathlib/Algebra/Star/Pointwise.lean @@ -5,7 +5,7 @@ Authors: Jireh Loreaux -/ import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.Star.Basic -import Mathlib.Data.Set.Finite +import Mathlib.Data.Set.Finite.Basic /-! # Pointwise star operation on sets diff --git a/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean b/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean index c0b03e905a1a6..db4c3e1b3959c 100644 --- a/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean +++ b/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean @@ -5,6 +5,7 @@ Authors: Joël Riou -/ import Mathlib.AlgebraicTopology.SimplicialObject import Mathlib.CategoryTheory.Limits.Shapes.Products +import Mathlib.Data.Fintype.Sigma /-! diff --git a/Mathlib/Analysis/Normed/Module/Dual.lean b/Mathlib/Analysis/Normed/Module/Dual.lean index 1a1c2f4219d0f..99cc8edec9dd3 100644 --- a/Mathlib/Analysis/Normed/Module/Dual.lean +++ b/Mathlib/Analysis/Normed/Module/Dual.lean @@ -3,10 +3,10 @@ Copyright (c) 2020 Heather Macbeth. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Heather Macbeth -/ +import Mathlib.Analysis.LocallyConvex.Polar import Mathlib.Analysis.NormedSpace.HahnBanach.Extension import Mathlib.Analysis.NormedSpace.RCLike -import Mathlib.Analysis.LocallyConvex.Polar -import Mathlib.Data.Set.Finite +import Mathlib.Data.Set.Finite.Lemmas /-! # The topological dual of a normed space diff --git a/Mathlib/CategoryTheory/Types.lean b/Mathlib/CategoryTheory/Types.lean index 5e029e2df1074..eaeb4a85402e2 100644 --- a/Mathlib/CategoryTheory/Types.lean +++ b/Mathlib/CategoryTheory/Types.lean @@ -5,8 +5,8 @@ Authors: Stephen Morgan, Kim Morrison, Johannes Hölzl -/ import Mathlib.CategoryTheory.EpiMono import Mathlib.CategoryTheory.Functor.FullyFaithful -import Mathlib.Tactic.PPWithUniv import Mathlib.Data.Set.Operations +import Mathlib.Tactic.PPWithUniv /-! # The category `Type`. diff --git a/Mathlib/Combinatorics/HalesJewett.lean b/Mathlib/Combinatorics/HalesJewett.lean index 76bb5d7812114..32e369960eca8 100644 --- a/Mathlib/Combinatorics/HalesJewett.lean +++ b/Mathlib/Combinatorics/HalesJewett.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Data.Fintype.Option import Mathlib.Data.Fintype.Shrink import Mathlib.Data.Fintype.Sum +import Mathlib.Data.Finite.Prod /-! # The Hales-Jewett theorem diff --git a/Mathlib/Combinatorics/Hall/Finite.lean b/Mathlib/Combinatorics/Hall/Finite.lean index a81d2cf78e3d0..658e17aae213d 100644 --- a/Mathlib/Combinatorics/Hall/Finite.lean +++ b/Mathlib/Combinatorics/Hall/Finite.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Alena Gusakov, Bhavik Mehta, Kyle Miller -/ import Mathlib.Data.Fintype.Basic -import Mathlib.Data.Set.Finite +import Mathlib.Data.Fintype.Powerset +import Mathlib.Data.Set.Finite.Basic /-! # Hall's Marriage Theorem for finite index types diff --git a/Mathlib/Combinatorics/SetFamily/Shatter.lean b/Mathlib/Combinatorics/SetFamily/Shatter.lean index 5e9757ad7d36a..27a76906f15d6 100644 --- a/Mathlib/Combinatorics/SetFamily/Shatter.lean +++ b/Mathlib/Combinatorics/SetFamily/Shatter.lean @@ -5,6 +5,7 @@ Authors: Yaël Dillies -/ import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Combinatorics.SetFamily.Compression.Down +import Mathlib.Data.Fintype.Powerset import Mathlib.Order.Interval.Finset.Nat /-! diff --git a/Mathlib/Combinatorics/SimpleGraph/Basic.lean b/Mathlib/Combinatorics/SimpleGraph/Basic.lean index b6c55a1bd0461..1da6a51b3e613 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Basic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Basic.lean @@ -4,8 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson, Jalex Stark, Kyle Miller, Alena Gusakov, Hunter Monroe -/ import Mathlib.Combinatorics.SimpleGraph.Init +import Mathlib.Data.Finite.Prod import Mathlib.Data.Rel -import Mathlib.Data.Set.Finite +import Mathlib.Data.Set.Finite.Basic import Mathlib.Data.Sym.Sym2 /-! diff --git a/Mathlib/Combinatorics/SimpleGraph/Clique.lean b/Mathlib/Combinatorics/SimpleGraph/Clique.lean index 31c644848116f..dd588e24a5611 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Clique.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Clique.lean @@ -6,6 +6,7 @@ Authors: Yaël Dillies, Bhavik Mehta import Mathlib.Combinatorics.SimpleGraph.Path import Mathlib.Combinatorics.SimpleGraph.Operations import Mathlib.Data.Finset.Pairwise +import Mathlib.Data.Fintype.Powerset import Mathlib.Data.Nat.Lattice /-! diff --git a/Mathlib/Combinatorics/SimpleGraph/Dart.lean b/Mathlib/Combinatorics/SimpleGraph/Dart.lean index 8919e4cea77bf..57d70b692b762 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Dart.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Dart.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller -/ import Mathlib.Combinatorics.SimpleGraph.Basic +import Mathlib.Data.Fintype.Sigma /-! # Darts in graphs diff --git a/Mathlib/Combinatorics/SimpleGraph/Finite.lean b/Mathlib/Combinatorics/SimpleGraph/Finite.lean index e0611556ee07e..cd15c4a2cff97 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Finite.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Finite.lean @@ -5,6 +5,7 @@ Authors: Aaron Anderson, Jalex Stark, Kyle Miller, Alena Gusakov -/ import Mathlib.Algebra.Order.Ring.Defs import Mathlib.Combinatorics.SimpleGraph.Basic +import Mathlib.Data.Finset.Max import Mathlib.Data.Sym.Card /-! diff --git a/Mathlib/Combinatorics/SimpleGraph/Maps.lean b/Mathlib/Combinatorics/SimpleGraph/Maps.lean index d18d062c11e50..4fc6a972b04d8 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Maps.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Maps.lean @@ -5,6 +5,7 @@ Authors: Hunter Monroe, Kyle Miller -/ import Mathlib.Combinatorics.SimpleGraph.Dart import Mathlib.Data.FunLike.Fintype +import Mathlib.Logic.Embedding.Set /-! # Maps between graphs diff --git a/Mathlib/Combinatorics/SimpleGraph/Matching.lean b/Mathlib/Combinatorics/SimpleGraph/Matching.lean index 3159fc24dc115..5df846e190c17 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Matching.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Matching.lean @@ -6,6 +6,7 @@ Authors: Alena Gusakov, Arthur Paulino, Kyle Miller, Pim Otte import Mathlib.Combinatorics.SimpleGraph.DegreeSum import Mathlib.Combinatorics.SimpleGraph.Connectivity.WalkCounting import Mathlib.Data.Fintype.Order +import Mathlib.Data.Set.Functor /-! # Matchings diff --git a/Mathlib/Combinatorics/SimpleGraph/StronglyRegular.lean b/Mathlib/Combinatorics/SimpleGraph/StronglyRegular.lean index 8b0a0607d5484..191e072fd1e6e 100644 --- a/Mathlib/Combinatorics/SimpleGraph/StronglyRegular.lean +++ b/Mathlib/Combinatorics/SimpleGraph/StronglyRegular.lean @@ -6,7 +6,6 @@ Authors: Alena Gusakov, Jeremy Tan import Mathlib.Combinatorics.Enumerative.DoubleCounting import Mathlib.Combinatorics.SimpleGraph.AdjMatrix import Mathlib.Combinatorics.SimpleGraph.Basic -import Mathlib.Data.Set.Finite /-! # Strongly regular graphs diff --git a/Mathlib/Data/DFinsupp/Defs.lean b/Mathlib/Data/DFinsupp/Defs.lean index 5289ba4147f35..c6780365cd08d 100644 --- a/Mathlib/Data/DFinsupp/Defs.lean +++ b/Mathlib/Data/DFinsupp/Defs.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Kenny Lau -/ -import Mathlib.Data.Set.Finite +import Mathlib.Data.Set.Finite.Basic /-! # Dependent functions with finite support diff --git a/Mathlib/Data/Finite/Defs.lean b/Mathlib/Data/Finite/Defs.lean index bc5e1a3774206..784e71e12e4f6 100644 --- a/Mathlib/Data/Finite/Defs.lean +++ b/Mathlib/Data/Finite/Defs.lean @@ -3,7 +3,10 @@ Copyright (c) 2022 Kyle Miller. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller -/ +import Mathlib.Data.Set.Operations import Mathlib.Logic.Equiv.Defs +import Mathlib.Tactic.Set +import Mathlib.Util.AssertExists /-! # Definition of the `Finite` typeclass @@ -28,18 +31,36 @@ instead. * `Finite α` denotes that `α` is a finite type. * `Infinite α` denotes that `α` is an infinite type. +* `Set.Finite : Set α → Prop` +* `Set.Infinite : Set α → Prop` +* `Set.toFinite` to prove `Set.Finite` for a `Set` from a `Finite` instance. ## Implementation notes +This file defines both the type-level `Finite` class and the set-level `Set.Finite` definition. + The definition of `Finite α` is not just `Nonempty (Fintype α)` since `Fintype` requires that `α : Type*`, and the definition in this module allows for `α : Sort*`. This means we can write the instance `Finite.prop`. +A finite set is defined to be a set whose coercion to a type has a `Finite` instance. + +There are two components to finiteness constructions. The first is `Fintype` instances for each +construction. This gives a way to actually compute a `Finset` that represents the set, and these +may be accessed using `set.toFinset`. This gets the `Finset` in the correct form, since otherwise +`Finset.univ : Finset s` is a `Finset` for the subtype for `s`. The second component is +"constructors" for `Set.Finite` that give proofs that `Fintype` instances exist classically given +other `Set.Finite` proofs. Unlike the `Fintype` instances, these *do not* use any decidability +instances since they do not compute anything. + ## Tags -finite, fintype +finite, fintype, finite sets -/ +assert_not_exists Finset +assert_not_exists MonoidWithZero +assert_not_exists OrderedRing universe u v @@ -140,3 +161,76 @@ protected theorem Infinite.false [Finite α] (_ : Infinite α) : False := @Infinite.not_finite α ‹_› ‹_› alias ⟨Finite.of_not_infinite, Finite.not_infinite⟩ := not_infinite_iff_finite + +section Set + +/-! +### Finite sets +-/ + +open Set Function + +variable {α : Type u} {β : Type v} + +namespace Set + +/-- A set is finite if the corresponding `Subtype` is finite, +i.e., if there exists a natural `n : ℕ` and an equivalence `s ≃ Fin n`. -/ +protected def Finite (s : Set α) : Prop := Finite s + +-- The `protected` attribute does not take effect within the same namespace block. +end Set + +namespace Set + +theorem finite_coe_iff {s : Set α} : Finite s ↔ s.Finite := .rfl + +/-- Constructor for `Set.Finite` using a `Finite` instance. -/ +theorem toFinite (s : Set α) [Finite s] : s.Finite := ‹_› + +/-- Projection of `Set.Finite` to its `Finite` instance. +This is intended to be used with dot notation. +See also `Set.Finite.Fintype` and `Set.Finite.nonempty_fintype`. -/ +protected theorem Finite.to_subtype {s : Set α} (h : s.Finite) : Finite s := h + +/-- A set is infinite if it is not finite. + +This is protected so that it does not conflict with global `Infinite`. -/ +protected def Infinite (s : Set α) : Prop := + ¬s.Finite + +@[simp] +theorem not_infinite {s : Set α} : ¬s.Infinite ↔ s.Finite := + not_not + +alias ⟨_, Finite.not_infinite⟩ := not_infinite + +attribute [simp] Finite.not_infinite + +/-- See also `finite_or_infinite`, `fintypeOrInfinite`. -/ +protected theorem finite_or_infinite (s : Set α) : s.Finite ∨ s.Infinite := + em _ + +protected theorem infinite_or_finite (s : Set α) : s.Infinite ∨ s.Finite := + em' _ + +end Set + +theorem Equiv.set_finite_iff {s : Set α} {t : Set β} (hst : s ≃ t) : s.Finite ↔ t.Finite := by + simp_rw [← Set.finite_coe_iff, hst.finite_iff] + +namespace Set + +/-! ### Infinite sets -/ + +variable {s t : Set α} + +theorem infinite_coe_iff {s : Set α} : Infinite s ↔ s.Infinite := + not_finite_iff_infinite.symm.trans finite_coe_iff.not + +-- Porting note: something weird happened here +alias ⟨_, Infinite.to_subtype⟩ := infinite_coe_iff + +end Set + +end Set diff --git a/Mathlib/Data/Finite/Prod.lean b/Mathlib/Data/Finite/Prod.lean index 4fd2053d88f57..42d9d71fa3edb 100644 --- a/Mathlib/Data/Finite/Prod.lean +++ b/Mathlib/Data/Finite/Prod.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Kyle Miller. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller -/ +import Mathlib.Data.Set.Finite.Basic import Mathlib.Data.Fintype.Prod import Mathlib.Data.Fintype.Vector @@ -10,6 +11,9 @@ import Mathlib.Data.Fintype.Vector # Finiteness of products -/ +assert_not_exists OrderedRing +assert_not_exists MonoidWithZero + open scoped Classical variable {α β : Type*} @@ -60,3 +64,171 @@ instance Equiv.finite_right {α β : Sort*} [Finite β] : Finite (α ≃ β) := instance Equiv.finite_left {α β : Sort*} [Finite α] : Finite (α ≃ β) := Finite.of_equiv _ ⟨Equiv.symm, Equiv.symm, Equiv.symm_symm, Equiv.symm_symm⟩ + +open Set Function + +variable {γ : Type*} + +namespace Set + +/-! ### Fintype instances + +Every instance here should have a corresponding `Set.Finite` constructor in the next section. +-/ + +section FintypeInstances + +instance fintypeProd (s : Set α) (t : Set β) [Fintype s] [Fintype t] : + Fintype (s ×ˢ t : Set (α × β)) := + Fintype.ofFinset (s.toFinset ×ˢ t.toFinset) <| by simp + +instance fintypeOffDiag [DecidableEq α] (s : Set α) [Fintype s] : Fintype s.offDiag := + Fintype.ofFinset s.toFinset.offDiag <| by simp + +/-- `image2 f s t` is `Fintype` if `s` and `t` are. -/ +instance fintypeImage2 [DecidableEq γ] (f : α → β → γ) (s : Set α) (t : Set β) [hs : Fintype s] + [ht : Fintype t] : Fintype (image2 f s t : Set γ) := by + rw [← image_prod] + apply Set.fintypeImage + +end FintypeInstances + +end Set + +/-! ### Finite instances + +There is seemingly some overlap between the following instances and the `Fintype` instances +in `Data.Set.Finite`. While every `Fintype` instance gives a `Finite` instance, those +instances that depend on `Fintype` or `Decidable` instances need an additional `Finite` instance +to be able to generally apply. + +Some set instances do not appear here since they are consequences of others, for example +`Subtype.Finite` for subsets of a finite type. +-/ + + +namespace Finite.Set + +open scoped Classical + +instance finite_prod (s : Set α) (t : Set β) [Finite s] [Finite t] : + Finite (s ×ˢ t : Set (α × β)) := + Finite.of_equiv _ (Equiv.Set.prod s t).symm + +instance finite_image2 (f : α → β → γ) (s : Set α) (t : Set β) [Finite s] [Finite t] : + Finite (image2 f s t : Set γ) := by + rw [← image_prod] + infer_instance + +end Finite.Set + +namespace Set + +/-! ### Constructors for `Set.Finite` + +Every constructor here should have a corresponding `Fintype` instance in the previous section +(or in the `Fintype` module). + +The implementation of these constructors ideally should be no more than `Set.toFinite`, +after possibly setting up some `Fintype` and classical `Decidable` instances. +-/ + + +section SetFiniteConstructors + +section Prod + +variable {s : Set α} {t : Set β} + +protected theorem Finite.prod (hs : s.Finite) (ht : t.Finite) : (s ×ˢ t : Set (α × β)).Finite := by + have := hs.to_subtype + have := ht.to_subtype + apply toFinite + +theorem Finite.of_prod_left (h : (s ×ˢ t : Set (α × β)).Finite) : t.Nonempty → s.Finite := + fun ⟨b, hb⟩ => (h.image Prod.fst).subset fun a ha => ⟨(a, b), ⟨ha, hb⟩, rfl⟩ + +theorem Finite.of_prod_right (h : (s ×ˢ t : Set (α × β)).Finite) : s.Nonempty → t.Finite := + fun ⟨a, ha⟩ => (h.image Prod.snd).subset fun b hb => ⟨(a, b), ⟨ha, hb⟩, rfl⟩ + +protected theorem Infinite.prod_left (hs : s.Infinite) (ht : t.Nonempty) : (s ×ˢ t).Infinite := + fun h => hs <| h.of_prod_left ht + +protected theorem Infinite.prod_right (ht : t.Infinite) (hs : s.Nonempty) : (s ×ˢ t).Infinite := + fun h => ht <| h.of_prod_right hs + +protected theorem infinite_prod : + (s ×ˢ t).Infinite ↔ s.Infinite ∧ t.Nonempty ∨ t.Infinite ∧ s.Nonempty := by + refine ⟨fun h => ?_, ?_⟩ + · simp_rw [Set.Infinite, @and_comm ¬_, ← Classical.not_imp] + by_contra! + exact h ((this.1 h.nonempty.snd).prod <| this.2 h.nonempty.fst) + · rintro (h | h) + · exact h.1.prod_left h.2 + · exact h.1.prod_right h.2 + +theorem finite_prod : (s ×ˢ t).Finite ↔ (s.Finite ∨ t = ∅) ∧ (t.Finite ∨ s = ∅) := by + simp only [← not_infinite, Set.infinite_prod, not_or, not_and_or, not_nonempty_iff_eq_empty] + +protected theorem Finite.offDiag {s : Set α} (hs : s.Finite) : s.offDiag.Finite := + (hs.prod hs).subset s.offDiag_subset_prod + +protected theorem Finite.image2 (f : α → β → γ) (hs : s.Finite) (ht : t.Finite) : + (image2 f s t).Finite := by + have := hs.to_subtype + have := ht.to_subtype + apply toFinite + +end Prod + +end SetFiniteConstructors + +/-! ### Properties -/ + +theorem Finite.toFinset_prod {s : Set α} {t : Set β} (hs : s.Finite) (ht : t.Finite) : + hs.toFinset ×ˢ ht.toFinset = (hs.prod ht).toFinset := + Finset.ext <| by simp + +theorem Finite.toFinset_offDiag {s : Set α} [DecidableEq α] (hs : s.Finite) : + hs.offDiag.toFinset = hs.toFinset.offDiag := + Finset.ext <| by simp + +theorem finite_image_fst_and_snd_iff {s : Set (α × β)} : + (Prod.fst '' s).Finite ∧ (Prod.snd '' s).Finite ↔ s.Finite := + ⟨fun h => (h.1.prod h.2).subset fun _ h => ⟨mem_image_of_mem _ h, mem_image_of_mem _ h⟩, + fun h => ⟨h.image _, h.image _⟩⟩ + +/-! ### Infinite sets -/ + +variable {s t : Set α} + +section Image2 + +variable {f : α → β → γ} {s : Set α} {t : Set β} {a : α} {b : β} + +protected theorem Infinite.image2_left (hs : s.Infinite) (hb : b ∈ t) + (hf : InjOn (fun a => f a b) s) : (image2 f s t).Infinite := + (hs.image hf).mono <| image_subset_image2_left hb + +protected theorem Infinite.image2_right (ht : t.Infinite) (ha : a ∈ s) (hf : InjOn (f a) t) : + (image2 f s t).Infinite := + (ht.image hf).mono <| image_subset_image2_right ha + +theorem infinite_image2 (hfs : ∀ b ∈ t, InjOn (fun a => f a b) s) (hft : ∀ a ∈ s, InjOn (f a) t) : + (image2 f s t).Infinite ↔ s.Infinite ∧ t.Nonempty ∨ t.Infinite ∧ s.Nonempty := by + refine ⟨fun h => Set.infinite_prod.1 ?_, ?_⟩ + · rw [← image_uncurry_prod] at h + exact h.of_image _ + · rintro (⟨hs, b, hb⟩ | ⟨ht, a, ha⟩) + · exact hs.image2_left hb (hfs _ hb) + · exact ht.image2_right ha (hft _ ha) + +lemma finite_image2 (hfs : ∀ b ∈ t, InjOn (f · b) s) (hft : ∀ a ∈ s, InjOn (f a) t) : + (image2 f s t).Finite ↔ s.Finite ∧ t.Finite ∨ s = ∅ ∨ t = ∅ := by + rw [← not_infinite, infinite_image2 hfs hft] + simp [not_or, -not_and, not_and_or, not_nonempty_iff_eq_empty] + aesop + +end Image2 + +end Set diff --git a/Mathlib/Data/Finset/Grade.lean b/Mathlib/Data/Finset/Grade.lean index 73c5f16bbbb90..813e23078c91b 100644 --- a/Mathlib/Data/Finset/Grade.lean +++ b/Mathlib/Data/Finset/Grade.lean @@ -4,9 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Data.Finset.Card -import Mathlib.Data.Set.Finite +import Mathlib.Data.Set.Finite.Basic import Mathlib.Order.Atoms import Mathlib.Order.Grade +import Mathlib.Order.Nat /-! # Finsets and multisets form a graded order diff --git a/Mathlib/Data/Finset/Interval.lean b/Mathlib/Data/Finset/Interval.lean index 3742d1f764991..2943c323cc06d 100644 --- a/Mathlib/Data/Finset/Interval.lean +++ b/Mathlib/Data/Finset/Interval.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Data.Finset.Grade +import Mathlib.Data.Finset.Powerset import Mathlib.Order.Interval.Finset.Basic /-! diff --git a/Mathlib/Data/Finset/NAry.lean b/Mathlib/Data/Finset/NAry.lean index a7d6db38c1b90..1ef61280f4ffc 100644 --- a/Mathlib/Data/Finset/NAry.lean +++ b/Mathlib/Data/Finset/NAry.lean @@ -3,8 +3,8 @@ 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.Data.Finset.Prod -import Mathlib.Data.Set.Finite +import Mathlib.Data.Finset.Lattice.Fold +import Mathlib.Data.Finite.Prod /-! # N-ary images of finsets diff --git a/Mathlib/Data/Finset/Preimage.lean b/Mathlib/Data/Finset/Preimage.lean index 1331695180e29..b302a82ddff61 100644 --- a/Mathlib/Data/Finset/Preimage.lean +++ b/Mathlib/Data/Finset/Preimage.lean @@ -3,8 +3,9 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ -import Mathlib.Data.Set.Finite +import Mathlib.Data.Finset.Sigma import Mathlib.Data.Finset.Sum +import Mathlib.Data.Set.Finite.Basic /-! # Preimage of a `Finset` under an injective map. diff --git a/Mathlib/Data/Finset/Slice.lean b/Mathlib/Data/Finset/Slice.lean index 577c05063d19a..2bb65d43b3036 100644 --- a/Mathlib/Data/Finset/Slice.lean +++ b/Mathlib/Data/Finset/Slice.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Bhavik Mehta, Alena Gusakov, Yaël Dillies -/ import Mathlib.Algebra.BigOperators.Group.Finset +import Mathlib.Data.Fintype.Powerset import Mathlib.Order.Antichain import Mathlib.Order.Interval.Finset.Nat diff --git a/Mathlib/Data/Finsupp/Defs.lean b/Mathlib/Data/Finsupp/Defs.lean index 8b6702f84c979..7ac40e922fd7b 100644 --- a/Mathlib/Data/Finsupp/Defs.lean +++ b/Mathlib/Data/Finsupp/Defs.lean @@ -5,7 +5,8 @@ Authors: Johannes Hölzl, Kim Morrison -/ import Mathlib.Algebra.Group.Indicator import Mathlib.Algebra.Group.Submonoid.Basic -import Mathlib.Data.Set.Finite +import Mathlib.Data.Finset.Max +import Mathlib.Data.Set.Finite.Basic /-! # Type of functions with finite support diff --git a/Mathlib/Data/Fintype/CardEmbedding.lean b/Mathlib/Data/Fintype/CardEmbedding.lean index da9e62d27dcf0..b67113680535a 100644 --- a/Mathlib/Data/Fintype/CardEmbedding.lean +++ b/Mathlib/Data/Fintype/CardEmbedding.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Rodriguez -/ import Mathlib.Data.Fintype.BigOperators +import Mathlib.Data.Set.Finite.Range import Mathlib.Logic.Equiv.Embedding /-! diff --git a/Mathlib/Data/Fintype/Fin.lean b/Mathlib/Data/Fintype/Fin.lean index 63559a8ab0c85..8744f61e2f3ae 100644 --- a/Mathlib/Data/Fintype/Fin.lean +++ b/Mathlib/Data/Fintype/Fin.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ import Mathlib.Order.Interval.Finset.Fin +import Mathlib.Data.Vector.Basic /-! # The structure of `Fintype (Fin n)` diff --git a/Mathlib/Data/Fintype/Order.lean b/Mathlib/Data/Fintype/Order.lean index 62b78c4e49d22..78516f9aea242 100644 --- a/Mathlib/Data/Fintype/Order.lean +++ b/Mathlib/Data/Fintype/Order.lean @@ -3,9 +3,11 @@ Copyright (c) 2021 Peter Nelson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Peter Nelson, Yaël Dillies -/ +import Mathlib.Data.Finset.Lattice.Fold import Mathlib.Data.Finset.Order +import Mathlib.Data.Set.Finite.Basic +import Mathlib.Data.Set.Finite.Range import Mathlib.Order.Atoms -import Mathlib.Data.Set.Finite import Mathlib.Order.Minimal /-! diff --git a/Mathlib/Data/Fintype/Pi.lean b/Mathlib/Data/Fintype/Pi.lean index 8fdb303b7f64b..2414aec167706 100644 --- a/Mathlib/Data/Fintype/Pi.lean +++ b/Mathlib/Data/Fintype/Pi.lean @@ -6,11 +6,15 @@ Authors: Mario Carneiro import Mathlib.Data.Finset.Card import Mathlib.Data.Finset.Pi import Mathlib.Data.Fintype.Basic +import Mathlib.Data.Set.Finite.Basic /-! # Fintype instances for pi types -/ +assert_not_exists OrderedRing +assert_not_exists MonoidWithZero + open Finset Function variable {α β : Type*} @@ -178,3 +182,43 @@ lemma piDiag_subset_piFinset [DecidableEq ι] [Fintype ι] : s.piDiag ι ⊆ Fintype.piFinset fun _ ↦ s := by simp [← piFinset_filter_const] end Finset + +namespace Set + +/-! ### Constructors for `Set.Finite` + +Every constructor here should have a corresponding `Fintype` instance in the previous section +(or in the `Fintype` module). + +The implementation of these constructors ideally should be no more than `Set.toFinite`, +after possibly setting up some `Fintype` and classical `Decidable` instances. +-/ + + +section SetFiniteConstructors + +section Pi +variable {ι : Type*} [Finite ι] {κ : ι → Type*} {t : ∀ i, Set (κ i)} + +/-- Finite product of finite sets is finite -/ +theorem Finite.pi (ht : ∀ i, (t i).Finite) : (pi univ t).Finite := by + cases nonempty_fintype ι + lift t to ∀ d, Finset (κ d) using ht + classical + rw [← Fintype.coe_piFinset] + apply Finset.finite_toSet + +/-- Finite product of finite sets is finite. Note this is a variant of `Set.Finite.pi` without the +extra `i ∈ univ` binder. -/ +lemma Finite.pi' (ht : ∀ i, (t i).Finite) : {f : ∀ i, κ i | ∀ i, f i ∈ t i}.Finite := by + simpa [Set.pi] using Finite.pi ht + +end Pi + +end SetFiniteConstructors + +theorem forall_finite_image_eval_iff {δ : Type*} [Finite δ] {κ : δ → Type*} {s : Set (∀ d, κ d)} : + (∀ d, (eval d '' s).Finite) ↔ s.Finite := + ⟨fun h => (Finite.pi h).subset <| subset_pi_eval_image _ _, fun h _ => h.image _⟩ + +end Set diff --git a/Mathlib/Data/MLList/BestFirst.lean b/Mathlib/Data/MLList/BestFirst.lean index 1eceab8e6500c..430ac91fa3323 100644 --- a/Mathlib/Data/MLList/BestFirst.lean +++ b/Mathlib/Data/MLList/BestFirst.lean @@ -5,8 +5,8 @@ Authors: Kim Morrison -/ import Batteries.Data.MLList.Basic import Mathlib.Data.Prod.Lex +import Mathlib.Data.Set.Finite.Range import Mathlib.Order.Estimator -import Mathlib.Data.Set.Finite /-! # Best first search diff --git a/Mathlib/Data/Matrix/Basic.lean b/Mathlib/Data/Matrix/Basic.lean index debd9e5f937d9..80fc79d3e62f4 100644 --- a/Mathlib/Data/Matrix/Basic.lean +++ b/Mathlib/Data/Matrix/Basic.lean @@ -6,6 +6,7 @@ Authors: Ellen Arlt, Blair Shi, Sean Leather, Mario Carneiro, Johan Commelin, Lu import Mathlib.Algebra.Algebra.Opposite import Mathlib.Algebra.Algebra.Pi import Mathlib.Algebra.BigOperators.RingEquiv +import Mathlib.Data.Finite.Prod import Mathlib.Data.Matrix.Mul import Mathlib.LinearAlgebra.Pi diff --git a/Mathlib/Data/Matroid/Basic.lean b/Mathlib/Data/Matroid/Basic.lean index 9031fedef544f..d3e21b40330bb 100644 --- a/Mathlib/Data/Matroid/Basic.lean +++ b/Mathlib/Data/Matroid/Basic.lean @@ -3,9 +3,10 @@ Copyright (c) 2023 Peter Nelson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Peter Nelson -/ +import Mathlib.Data.Finite.Prod +import Mathlib.Data.Matroid.Init import Mathlib.Data.Set.Card import Mathlib.Order.Minimal -import Mathlib.Data.Matroid.Init /-! # Matroids diff --git a/Mathlib/Data/Nat/PrimeFin.lean b/Mathlib/Data/Nat/PrimeFin.lean index ecb04ce9d0c0a..a9054c2d80174 100644 --- a/Mathlib/Data/Nat/PrimeFin.lean +++ b/Mathlib/Data/Nat/PrimeFin.lean @@ -5,8 +5,8 @@ Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Mathlib.Data.Countable.Defs import Mathlib.Data.Nat.Factors -import Mathlib.Data.Set.Finite import Mathlib.Data.Nat.Prime.Infinite +import Mathlib.Data.Set.Finite.Lattice /-! # Prime numbers diff --git a/Mathlib/Data/Set/Countable.lean b/Mathlib/Data/Set/Countable.lean index 7540d71c1c4ad..b22bd0f331322 100644 --- a/Mathlib/Data/Set/Countable.lean +++ b/Mathlib/Data/Set/Countable.lean @@ -3,10 +3,10 @@ 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 -/ -import Mathlib.Data.Set.Finite import Mathlib.Data.Countable.Basic -import Mathlib.Logic.Equiv.List +import Mathlib.Data.Set.Finite.Basic import Mathlib.Data.Set.Subsingleton +import Mathlib.Logic.Equiv.List /-! # Countable sets diff --git a/Mathlib/Data/Set/Finite.lean b/Mathlib/Data/Set/Finite/Basic.lean similarity index 57% rename from Mathlib/Data/Set/Finite.lean rename to Mathlib/Data/Set/Finite/Basic.lean index 30f991e729f12..800c3708d2d18 100644 --- a/Mathlib/Data/Set/Finite.lean +++ b/Mathlib/Data/Set/Finite/Basic.lean @@ -3,27 +3,21 @@ 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.Max -import Mathlib.Data.Set.Functor -import Mathlib.Data.Set.Lattice -import Mathlib.Data.Finite.Powerset -import Mathlib.Data.Finite.Prod -import Mathlib.Data.Finite.Sigma -import Mathlib.Data.Finite.Vector -import Mathlib.Logic.Embedding.Set +import Mathlib.Data.Finite.Defs +import Mathlib.Data.Fintype.Card /-! # Finite sets -This file defines predicates for finite and infinite sets and provides -`Fintype` instances for many set constructions. It also proves basic facts -about finite sets and gives ways to manipulate `Set.Finite` expressions. +This file provides `Fintype` instances for many set constructions. It also proves basic facts about +finite sets and gives ways to manipulate `Set.Finite` expressions. + +Note that the instances in this file are selected somewhat arbitrarily on the basis of them not +needing any imports beyond `Data.Fintype.Card` (which is required by `Finite.ofFinset`); they can +certainly be organized better. ## Main definitions -* `Set.Finite : Set α → Prop` -* `Set.Infinite : Set α → Prop` -* `Set.toFinite` to prove `Set.Finite` for a `Set` from a `Finite` instance. * `Set.Finite.toFinset` to noncomputably produce a `Finset` from a `Set.Finite` proof. (See `Set.toFinset` for a computable version.) @@ -55,34 +49,15 @@ variable {α : Type u} {β : Type v} {ι : Sort w} {γ : Type x} namespace Set -/-- A set is finite if the corresponding `Subtype` is finite, -i.e., if there exists a natural `n : ℕ` and an equivalence `s ≃ Fin n`. -/ -protected def Finite (s : Set α) : Prop := Finite s - --- The `protected` attribute does not take effect within the same namespace block. -end Set - -namespace Set - theorem finite_def {s : Set α} : s.Finite ↔ Nonempty (Fintype s) := finite_iff_nonempty_fintype s protected alias ⟨Finite.nonempty_fintype, _⟩ := finite_def -theorem finite_coe_iff {s : Set α} : Finite s ↔ s.Finite := .rfl - -/-- Constructor for `Set.Finite` using a `Finite` instance. -/ -theorem toFinite (s : Set α) [Finite s] : s.Finite := ‹_› - /-- Construct a `Finite` instance for a `Set` from a `Finset` with the same elements. -/ protected theorem Finite.ofFinset {p : Set α} (s : Finset α) (H : ∀ x, x ∈ s ↔ x ∈ p) : p.Finite := have := Fintype.ofFinset s H; p.toFinite -/-- Projection of `Set.Finite` to its `Finite` instance. -This is intended to be used with dot notation. -See also `Set.Finite.Fintype` and `Set.Finite.nonempty_fintype`. -/ -protected theorem Finite.to_subtype {s : Set α} (h : s.Finite) : Finite s := h - /-- A finite set coerced to a type is a `Fintype`. This is the `Fintype` projection for a `Set.Finite`. @@ -118,27 +93,6 @@ theorem Finite.exists_finset_coe {s : Set α} (h : s.Finite) : ∃ s' : Finset /-- Finite sets can be lifted to finsets. -/ instance : CanLift (Set α) (Finset α) (↑) Set.Finite where prf _ hs := hs.exists_finset_coe -/-- A set is infinite if it is not finite. - -This is protected so that it does not conflict with global `Infinite`. -/ -protected def Infinite (s : Set α) : Prop := - ¬s.Finite - -@[simp] -theorem not_infinite {s : Set α} : ¬s.Infinite ↔ s.Finite := - not_not - -alias ⟨_, Finite.not_infinite⟩ := not_infinite - -attribute [simp] Finite.not_infinite - -/-- See also `finite_or_infinite`, `fintypeOrInfinite`. -/ -protected theorem finite_or_infinite (s : Set α) : s.Finite ∨ s.Infinite := - em _ - -protected theorem infinite_or_finite (s : Set α) : s.Infinite ∨ s.Finite := - em' _ - /-! ### Basic properties of `Set.Finite.toFinset` -/ @@ -322,15 +276,6 @@ instance fintypeDiffLeft (s t : Set α) [Fintype s] [DecidablePred (· ∈ t)] : Fintype (s \ t : Set α) := Set.fintypeSep s (· ∈ tᶜ) -instance fintypeiUnion [DecidableEq α] [Fintype (PLift ι)] (f : ι → Set α) [∀ i, Fintype (f i)] : - Fintype (⋃ i, f i) := - Fintype.ofFinset (Finset.univ.biUnion fun i : PLift ι => (f i.down).toFinset) <| by simp - -instance fintypesUnion [DecidableEq α] {s : Set (Set α)} [Fintype s] - [H : ∀ t : s, Fintype (t : Set α)] : Fintype (⋃₀ s) := by - rw [sUnion_eq_iUnion] - exact @Set.fintypeiUnion _ _ _ _ _ H - /-- A union of sets with `Fintype` structure over a set with `Fintype` structure has a `Fintype` structure. -/ def fintypeBiUnion [DecidableEq α] {ι : Type*} (s : Set ι) [Fintype s] (t : ι → Set α) @@ -342,37 +287,12 @@ instance fintypeBiUnion' [DecidableEq α] {ι : Type*} (s : Set ι) [Fintype s] [∀ i, Fintype (t i)] : Fintype (⋃ x ∈ s, t x) := Fintype.ofFinset (s.toFinset.biUnion fun x => (t x).toFinset) <| by simp -lemma toFinset_iUnion [Fintype β] [DecidableEq α] (f : β → Set α) - [∀ w, Fintype (f w)] : - Set.toFinset (⋃ (x : β), f x) = - Finset.biUnion (Finset.univ : Finset β) (fun x => (f x).toFinset) := by - ext v - simp only [mem_toFinset, mem_iUnion, Finset.mem_biUnion, Finset.mem_univ, true_and] - -section monad -attribute [local instance] Set.monad - -/-- If `s : Set α` is a set with `Fintype` instance and `f : α → Set β` is a function such that -each `f a`, `a ∈ s`, has a `Fintype` structure, then `s >>= f` has a `Fintype` structure. -/ -def fintypeBind {α β} [DecidableEq β] (s : Set α) [Fintype s] (f : α → Set β) - (H : ∀ a ∈ s, Fintype (f a)) : Fintype (s >>= f) := - Set.fintypeBiUnion s f H - -instance fintypeBind' {α β} [DecidableEq β] (s : Set α) [Fintype s] (f : α → Set β) - [∀ a, Fintype (f a)] : Fintype (s >>= f) := - Set.fintypeBiUnion' s f - -end monad - instance fintypeEmpty : Fintype (∅ : Set α) := Fintype.ofFinset ∅ <| by simp instance fintypeSingleton (a : α) : Fintype ({a} : Set α) := Fintype.ofFinset {a} <| by simp -instance fintypePure : ∀ a : α, Fintype (pure a : Set α) := - Set.fintypeSingleton - /-- A `Fintype` instance for inserting an element into a `Set` using the corresponding `insert` function on `Finset`. This requires `DecidableEq α`. There is also `Set.fintypeInsert'` when `a ∈ s` is decidable. -/ @@ -416,9 +336,6 @@ def fintypeOfFintypeImage (s : Set α) {f : α → β} {g} (I : IsPartialInv f g suffices (∃ x, x ∈ s ∧ g (f x) = some a) ↔ a ∈ s by simpa [and_comm, and_left_comm, and_assoc] simp [I _, (injective_of_isPartialInv I).eq_iff] -instance fintypeRange [DecidableEq α] (f : ι → α) [Fintype (PLift ι)] : Fintype (range f) := - Fintype.ofFinset (Finset.univ.image <| f ∘ PLift.down) <| by simp - instance fintypeMap {α β} [DecidableEq β] : ∀ (s : Set α) (f : α → β) [Fintype s], Fintype (f <$> s) := Set.fintypeImage @@ -434,28 +351,6 @@ in `Mathlib/Order/LocallyFinite.lean`. -/ def Nat.fintypeIio (n : ℕ) : Fintype (Iio n) := Set.fintypeLTNat n -instance fintypeProd (s : Set α) (t : Set β) [Fintype s] [Fintype t] : - Fintype (s ×ˢ t : Set (α × β)) := - Fintype.ofFinset (s.toFinset ×ˢ t.toFinset) <| by simp - -instance fintypeOffDiag [DecidableEq α] (s : Set α) [Fintype s] : Fintype s.offDiag := - Fintype.ofFinset s.toFinset.offDiag <| by simp - -/-- `image2 f s t` is `Fintype` if `s` and `t` are. -/ -instance fintypeImage2 [DecidableEq γ] (f : α → β → γ) (s : Set α) (t : Set β) [hs : Fintype s] - [ht : Fintype t] : Fintype (image2 f s t : Set γ) := by - rw [← image_prod] - apply Set.fintypeImage - -instance fintypeSeq [DecidableEq β] (f : Set (α → β)) (s : Set α) [Fintype f] [Fintype s] : - Fintype (f.seq s) := by - rw [seq_def] - apply Set.fintypeBiUnion' - -instance fintypeSeq' {α β : Type u} [DecidableEq β] (f : Set (α → β)) (s : Set α) [Fintype f] - [Fintype s] : Fintype (f <*> s) := - Set.fintypeSeq f s - instance fintypeMemFinset (s : Finset α) : Fintype { a | a ∈ s } := Finset.fintypeCoeSort s @@ -463,9 +358,6 @@ end FintypeInstances end Set -theorem Equiv.set_finite_iff {s : Set α} {t : Set β} (hst : s ≃ t) : s.Finite ↔ t.Finite := by - simp_rw [← Set.finite_coe_iff, hst.finite_iff] - /-! ### Finset -/ namespace Finset @@ -546,40 +438,6 @@ instance finite_inter_of_left (s t : Set α) [Finite s] : Finite (s ∩ t : Set instance finite_diff (s t : Set α) [Finite s] : Finite (s \ t : Set α) := Finite.Set.subset s diff_subset -instance finite_range (f : ι → α) [Finite ι] : Finite (range f) := by - haveI := Fintype.ofFinite (PLift ι) - infer_instance - -instance finite_iUnion [Finite ι] (f : ι → Set α) [∀ i, Finite (f i)] : Finite (⋃ i, f i) := by - rw [iUnion_eq_range_psigma] - apply Set.finite_range - -instance finite_sUnion {s : Set (Set α)} [Finite s] [H : ∀ t : s, Finite (t : Set α)] : - Finite (⋃₀ s) := by - rw [sUnion_eq_iUnion] - exact @Finite.Set.finite_iUnion _ _ _ _ H - -theorem finite_biUnion {ι : Type*} (s : Set ι) [Finite s] (t : ι → Set α) - (H : ∀ i ∈ s, Finite (t i)) : Finite (⋃ x ∈ s, t x) := by - rw [biUnion_eq_iUnion] - haveI : ∀ i : s, Finite (t i) := fun i => H i i.property - infer_instance - -instance finite_biUnion' {ι : Type*} (s : Set ι) [Finite s] (t : ι → Set α) [∀ i, Finite (t i)] : - Finite (⋃ x ∈ s, t x) := - finite_biUnion s t fun _ _ => inferInstance - -/-- Example: `Finite (⋃ (i < n), f i)` where `f : ℕ → Set α` and `[∀ i, Finite (f i)]` -(when given instances from `Order.Interval.Finset.Nat`). --/ -instance finite_biUnion'' {ι : Type*} (p : ι → Prop) [h : Finite { x | p x }] (t : ι → Set α) - [∀ i, Finite (t i)] : Finite (⋃ (x) (_ : p x), t x) := - @Finite.Set.finite_biUnion' _ _ (setOf p) h t _ - -instance finite_iInter {ι : Sort*} [Nonempty ι] (t : ι → Set α) [∀ i, Finite (t i)] : - Finite (⋂ i, t i) := - Finite.Set.subset (t <| Classical.arbitrary ι) (iInter_subset _ _) - instance finite_insert (a : α) (s : Set α) [Finite s] : Finite (insert a s : Set α) := Finite.Set.finite_union {a} s @@ -587,23 +445,6 @@ instance finite_image (s : Set α) (f : α → β) [Finite s] : Finite (f '' s) cases nonempty_fintype s infer_instance -instance finite_replacement [Finite α] (f : α → β) : - Finite {f x | x : α} := - Finite.Set.finite_range f - -instance finite_prod (s : Set α) (t : Set β) [Finite s] [Finite t] : - Finite (s ×ˢ t : Set (α × β)) := - Finite.of_equiv _ (Equiv.Set.prod s t).symm - -instance finite_image2 (f : α → β → γ) (s : Set α) (t : Set β) [Finite s] [Finite t] : - Finite (image2 f s t : Set γ) := by - rw [← image_prod] - infer_instance - -instance finite_seq (f : Set (α → β)) (s : Set α) [Finite f] [Finite s] : Finite (f.seq s) := by - rw [seq_def] - infer_instance - end Finite.Set namespace Set @@ -670,49 +511,6 @@ theorem Finite.diff {s : Set α} (hs : s.Finite) (t : Set α) : (s \ t).Finite : theorem Finite.of_diff {s t : Set α} (hd : (s \ t).Finite) (ht : t.Finite) : s.Finite := (hd.union ht).subset <| subset_diff_union _ _ -theorem finite_iUnion [Finite ι] {f : ι → Set α} (H : ∀ i, (f i).Finite) : (⋃ i, f i).Finite := - haveI := fun i => (H i).to_subtype - toFinite _ - -/-- Dependent version of `Finite.biUnion`. -/ -theorem Finite.biUnion' {ι} {s : Set ι} (hs : s.Finite) {t : ∀ i ∈ s, Set α} - (ht : ∀ i (hi : i ∈ s), (t i hi).Finite) : (⋃ i ∈ s, t i ‹_›).Finite := by - have := hs.to_subtype - rw [biUnion_eq_iUnion] - apply finite_iUnion fun i : s => ht i.1 i.2 - -theorem Finite.biUnion {ι} {s : Set ι} (hs : s.Finite) {t : ι → Set α} - (ht : ∀ i ∈ s, (t i).Finite) : (⋃ i ∈ s, t i).Finite := - hs.biUnion' ht - -theorem Finite.sUnion {s : Set (Set α)} (hs : s.Finite) (H : ∀ t ∈ s, Set.Finite t) : - (⋃₀ s).Finite := by - simpa only [sUnion_eq_biUnion] using hs.biUnion H - -theorem Finite.sInter {α : Type*} {s : Set (Set α)} {t : Set α} (ht : t ∈ s) (hf : t.Finite) : - (⋂₀ s).Finite := - hf.subset (sInter_subset_of_mem ht) - -/-- If sets `s i` are finite for all `i` from a finite set `t` and are empty for `i ∉ t`, then the -union `⋃ i, s i` is a finite set. -/ -theorem Finite.iUnion {ι : Type*} {s : ι → Set α} {t : Set ι} (ht : t.Finite) - (hs : ∀ i ∈ t, (s i).Finite) (he : ∀ i, i ∉ t → s i = ∅) : (⋃ i, s i).Finite := by - suffices ⋃ i, s i ⊆ ⋃ i ∈ t, s i by exact (ht.biUnion hs).subset this - refine iUnion_subset fun i x hx => ?_ - by_cases hi : i ∈ t - · exact mem_biUnion hi hx - · rw [he i hi, mem_empty_iff_false] at hx - contradiction - -section monad -attribute [local instance] Set.monad - -theorem Finite.bind {α β} {s : Set α} {f : α → Set β} (h : s.Finite) (hf : ∀ a ∈ s, (f a).Finite) : - (s >>= f).Finite := - h.biUnion hf - -end monad - @[simp] theorem finite_empty : (∅ : Set α).Finite := toFinite _ @@ -726,9 +524,6 @@ protected theorem Infinite.nonempty {s : Set α} (h : s.Infinite) : s.Nonempty : theorem finite_singleton (a : α) : ({a} : Set α).Finite := toFinite _ -theorem finite_pure (a : α) : (pure a : Set α).Finite := - toFinite _ - @[simp] protected theorem Finite.insert (a : α) {s : Set α} (hs : s.Finite) : (insert a s).Finite := (finite_singleton a).union hs @@ -737,17 +532,9 @@ theorem Finite.image {s : Set α} (f : α → β) (hs : s.Finite) : (f '' s).Fin have := hs.to_subtype apply toFinite -theorem finite_range (f : ι → α) [Finite ι] : (range f).Finite := - toFinite _ - lemma Finite.of_surjOn {s : Set α} {t : Set β} (f : α → β) (hf : SurjOn f s t) (hs : s.Finite) : t.Finite := (hs.image _).subset hf -theorem Finite.dependent_image {s : Set α} (hs : s.Finite) (F : ∀ i ∈ s, β) : - {y : β | ∃ x hx, F x hx = y}.Finite := by - have := hs.to_subtype - simpa [range] using finite_range fun x : s => F x x.2 - theorem Finite.map {α β} {s : Set α} : ∀ f : α → β, s.Finite → (f <$> s).Finite := Finite.image @@ -769,11 +556,6 @@ theorem Finite.of_preimage (h : (f ⁻¹' s).Finite) (hf : Surjective f) : s.Fin theorem Finite.preimage (I : Set.InjOn f (f ⁻¹' s)) (h : s.Finite) : (f ⁻¹' s).Finite := (h.subset (image_preimage_subset f s)).of_finite_image I -theorem Finite.preimage' (h : s.Finite) (hf : ∀ b ∈ s, (f ⁻¹' {b}).Finite) : - (f ⁻¹' s).Finite := by - rw [← Set.biUnion_preimage_singleton] - exact Set.Finite.biUnion h hf - protected lemma Infinite.preimage (hs : s.Infinite) (hf : s ⊆ range f) : (f ⁻¹' s).Infinite := fun h ↦ hs <| finite_of_finite_preimage h hf @@ -809,59 +591,6 @@ theorem Finite.injOn_iff_bijOn_of_mapsTo (hs : s.Finite) (hm : MapsTo f s s) : end MapsTo -section Prod - -variable {s : Set α} {t : Set β} - -protected theorem Finite.prod (hs : s.Finite) (ht : t.Finite) : (s ×ˢ t : Set (α × β)).Finite := by - have := hs.to_subtype - have := ht.to_subtype - apply toFinite - -theorem Finite.of_prod_left (h : (s ×ˢ t : Set (α × β)).Finite) : t.Nonempty → s.Finite := - fun ⟨b, hb⟩ => (h.image Prod.fst).subset fun a ha => ⟨(a, b), ⟨ha, hb⟩, rfl⟩ - -theorem Finite.of_prod_right (h : (s ×ˢ t : Set (α × β)).Finite) : s.Nonempty → t.Finite := - fun ⟨a, ha⟩ => (h.image Prod.snd).subset fun b hb => ⟨(a, b), ⟨ha, hb⟩, rfl⟩ - -protected theorem Infinite.prod_left (hs : s.Infinite) (ht : t.Nonempty) : (s ×ˢ t).Infinite := - fun h => hs <| h.of_prod_left ht - -protected theorem Infinite.prod_right (ht : t.Infinite) (hs : s.Nonempty) : (s ×ˢ t).Infinite := - fun h => ht <| h.of_prod_right hs - -protected theorem infinite_prod : - (s ×ˢ t).Infinite ↔ s.Infinite ∧ t.Nonempty ∨ t.Infinite ∧ s.Nonempty := by - refine ⟨fun h => ?_, ?_⟩ - · simp_rw [Set.Infinite, @and_comm ¬_, ← Classical.not_imp] - by_contra! - exact h ((this.1 h.nonempty.snd).prod <| this.2 h.nonempty.fst) - · rintro (h | h) - · exact h.1.prod_left h.2 - · exact h.1.prod_right h.2 - -theorem finite_prod : (s ×ˢ t).Finite ↔ (s.Finite ∨ t = ∅) ∧ (t.Finite ∨ s = ∅) := by - simp only [← not_infinite, Set.infinite_prod, not_or, not_and_or, not_nonempty_iff_eq_empty] - -protected theorem Finite.offDiag {s : Set α} (hs : s.Finite) : s.offDiag.Finite := - (hs.prod hs).subset s.offDiag_subset_prod - -protected theorem Finite.image2 (f : α → β → γ) (hs : s.Finite) (ht : t.Finite) : - (image2 f s t).Finite := by - have := hs.to_subtype - have := ht.to_subtype - apply toFinite - -end Prod - -theorem Finite.seq {f : Set (α → β)} {s : Set α} (hf : f.Finite) (hs : s.Finite) : - (f.seq s).Finite := - hf.image2 _ hs - -theorem Finite.seq' {α β : Type u} {f : Set (α → β)} {s : Set α} (hf : f.Finite) (hs : s.Finite) : - (f <*> s).Finite := - hf.seq hs - theorem finite_mem_finset (s : Finset α) : { a | a ∈ s }.Finite := toFinite _ @@ -881,16 +610,6 @@ theorem exists_finite_iff_finset {p : Set α → Prop} : ⟨fun ⟨_, hs, hps⟩ => ⟨hs.toFinset, hs.coe_toFinset.symm ▸ hps⟩, fun ⟨s, hs⟩ => ⟨s, s.finite_toSet, hs⟩⟩ -/-- There are finitely many subsets of a given finite set -/ -theorem Finite.finite_subsets {α : Type u} {a : Set α} (h : a.Finite) : { b | b ⊆ a }.Finite := by - convert ((Finset.powerset h.toFinset).map Finset.coeEmb.1).finite_toSet - ext s - simpa [← @exists_finite_iff_finset α fun t => t ⊆ a ∧ t = s, Finite.subset_toFinset, - ← and_assoc, Finset.coeEmb] using h.subset - -protected theorem Finite.powerset {s : Set α} (h : s.Finite) : (𝒫 s).Finite := - h.finite_subsets - theorem exists_subset_image_finite_and {f : α → β} {s : Set α} {p : Set β → Prop} : (∃ t ⊆ f '' s, t.Finite ∧ p t) ↔ ∃ t ⊆ s, t.Finite ∧ p (f '' t) := by classical @@ -898,30 +617,6 @@ theorem exists_subset_image_finite_and {f : α → β} {s : Set α} {p : Set β Finset.subset_set_image_iff] aesop -section Pi -variable {ι : Type*} [Finite ι] {κ : ι → Type*} {t : ∀ i, Set (κ i)} - -/-- Finite product of finite sets is finite -/ -theorem Finite.pi (ht : ∀ i, (t i).Finite) : (pi univ t).Finite := by - cases nonempty_fintype ι - lift t to ∀ d, Finset (κ d) using ht - classical - rw [← Fintype.coe_piFinset] - apply Finset.finite_toSet - -/-- Finite product of finite sets is finite. Note this is a variant of `Set.Finite.pi` without the -extra `i ∈ univ` binder. -/ -lemma Finite.pi' (ht : ∀ i, (t i).Finite) : {f : ∀ i, κ i | ∀ i, f i ∈ t i}.Finite := by - simpa [Set.pi] using Finite.pi ht - -end Pi - -/-- A finite union of finsets is finite. -/ -theorem union_finset_finite_of_range_finite (f : α → Finset β) (h : (range f).Finite) : - (⋃ a, (f a : Set β)).Finite := by - rw [← biUnion_range] - exact h.biUnion fun y _ => y.finite_toSet - theorem finite_range_ite {p : α → Prop} [DecidablePred p] {f g : α → β} (hf : (range f).Finite) (hg : (range g).Finite) : (range fun x => if p x then f x else g x).Finite := (hf.union hg).subset range_ite_subset @@ -961,63 +656,11 @@ theorem Finite.toFinset_insert' [DecidableEq α] {a : α} {s : Set α} (hs : s.F (hs.insert a).toFinset = insert a hs.toFinset := Finite.toFinset_insert _ -theorem Finite.toFinset_prod {s : Set α} {t : Set β} (hs : s.Finite) (ht : t.Finite) : - hs.toFinset ×ˢ ht.toFinset = (hs.prod ht).toFinset := - Finset.ext <| by simp - -theorem Finite.toFinset_offDiag {s : Set α} [DecidableEq α] (hs : s.Finite) : - hs.offDiag.toFinset = hs.toFinset.offDiag := - Finset.ext <| by simp - -theorem Finite.fin_embedding {s : Set α} (h : s.Finite) : - ∃ (n : ℕ) (f : Fin n ↪ α), range f = s := - ⟨_, (Fintype.equivFin (h.toFinset : Set α)).symm.asEmbedding, by - simp only [Finset.coe_sort_coe, Equiv.asEmbedding_range, Finite.coe_toFinset, setOf_mem_eq]⟩ - -theorem Finite.fin_param {s : Set α} (h : s.Finite) : - ∃ (n : ℕ) (f : Fin n → α), Injective f ∧ range f = s := - let ⟨n, f, hf⟩ := h.fin_embedding - ⟨n, f, f.injective, hf⟩ - theorem finite_option {s : Set (Option α)} : s.Finite ↔ { x : α | some x ∈ s }.Finite := ⟨fun h => h.preimage_embedding Embedding.some, fun h => ((h.image some).insert none).subset fun x => x.casesOn (fun _ => Or.inl rfl) fun _ hx => Or.inr <| mem_image_of_mem _ hx⟩ -theorem finite_image_fst_and_snd_iff {s : Set (α × β)} : - (Prod.fst '' s).Finite ∧ (Prod.snd '' s).Finite ↔ s.Finite := - ⟨fun h => (h.1.prod h.2).subset fun _ h => ⟨mem_image_of_mem _ h, mem_image_of_mem _ h⟩, - fun h => ⟨h.image _, h.image _⟩⟩ - -theorem forall_finite_image_eval_iff {δ : Type*} [Finite δ] {κ : δ → Type*} {s : Set (∀ d, κ d)} : - (∀ d, (eval d '' s).Finite) ↔ s.Finite := - ⟨fun h => (Finite.pi h).subset <| subset_pi_eval_image _ _, fun h _ => h.image _⟩ - -theorem finite_subset_iUnion {s : Set α} (hs : s.Finite) {ι} {t : ι → Set α} (h : s ⊆ ⋃ i, t i) : - ∃ I : Set ι, I.Finite ∧ s ⊆ ⋃ i ∈ I, t i := by - have := hs.to_subtype - choose f hf using show ∀ x : s, ∃ i, x.1 ∈ t i by simpa [subset_def] using h - refine ⟨range f, finite_range f, fun x hx => ?_⟩ - rw [biUnion_range, mem_iUnion] - exact ⟨⟨x, hx⟩, hf _⟩ - -theorem eq_finite_iUnion_of_finite_subset_iUnion {ι} {s : ι → Set α} {t : Set α} (tfin : t.Finite) - (h : t ⊆ ⋃ i, s i) : - ∃ I : Set ι, - I.Finite ∧ - ∃ σ : { i | i ∈ I } → Set α, (∀ i, (σ i).Finite) ∧ (∀ i, σ i ⊆ s i) ∧ t = ⋃ i, σ i := - let ⟨I, Ifin, hI⟩ := finite_subset_iUnion tfin h - ⟨I, Ifin, fun x => s x ∩ t, fun _ => tfin.subset inter_subset_right, fun _ => - inter_subset_left, by - ext x - rw [mem_iUnion] - constructor - · intro x_in - rcases mem_iUnion.mp (hI x_in) with ⟨i, _, ⟨hi, rfl⟩, H⟩ - exact ⟨⟨i, hi⟩, ⟨H, x_in⟩⟩ - · rintro ⟨i, -, H⟩ - exact H⟩ - @[elab_as_elim] theorem Finite.induction_on {C : Set α → Prop} {s : Set α} (h : s.Finite) (H0 : C ∅) (H1 : ∀ {a s}, a ∉ s → Set.Finite s → C s → C (insert a s)) : C s := by @@ -1044,23 +687,6 @@ theorem Finite.dinduction_on {C : ∀ s : Set α, s.Finite → Prop} (s : Set α Finite.induction_on h (fun _ => H0) fun has hs ih _ => H1 has hs (ih _) this h -/-- Induction up to a finite set `S`. -/ -theorem Finite.induction_to {C : Set α → Prop} {S : Set α} (h : S.Finite) - (S0 : Set α) (hS0 : S0 ⊆ S) (H0 : C S0) (H1 : ∀ s ⊂ S, C s → ∃ a ∈ S \ s, C (insert a s)) : - C S := by - have : Finite S := Finite.to_subtype h - have : Finite {T : Set α // T ⊆ S} := Finite.of_equiv (Set S) (Equiv.Set.powerset S).symm - rw [← Subtype.coe_mk (p := (· ⊆ S)) _ le_rfl] - rw [← Subtype.coe_mk (p := (· ⊆ S)) _ hS0] at H0 - refine Finite.to_wellFoundedGT.wf.induction_bot' (fun s hs hs' ↦ ?_) H0 - obtain ⟨a, ⟨ha1, ha2⟩, ha'⟩ := H1 s (ssubset_of_ne_of_subset hs s.2) hs' - exact ⟨⟨insert a s.1, insert_subset ha1 s.2⟩, Set.ssubset_insert ha2, ha'⟩ - -/-- Induction up to `univ`. -/ -theorem Finite.induction_to_univ [Finite α] {C : Set α → Prop} (S0 : Set α) - (H0 : C S0) (H1 : ∀ S ≠ univ, C S → ∃ a ∉ S, C (insert a S)) : C univ := - finite_univ.induction_to S0 (subset_univ S0) H0 (by simpa [ssubset_univ_iff]) - section attribute [local instance] Nat.fintypeIio @@ -1156,12 +782,6 @@ theorem infinite_univ_iff : (@univ α).Infinite ↔ Infinite α := by theorem infinite_univ [h : Infinite α] : (@univ α).Infinite := infinite_univ_iff.2 h -theorem infinite_coe_iff {s : Set α} : Infinite s ↔ s.Infinite := - not_finite_iff_infinite.symm.trans finite_coe_iff.not - --- Porting note: something weird happened here -alias ⟨_, Infinite.to_subtype⟩ := infinite_coe_iff - lemma Infinite.exists_not_mem_finite (hs : s.Infinite) (ht : t.Finite) : ∃ a, a ∈ s ∧ a ∉ t := by by_contra! h; exact hs <| ht.subset h @@ -1212,35 +832,6 @@ theorem infinite_range_iff {f : α → β} (hi : Injective f) : protected alias ⟨_, Infinite.image⟩ := infinite_image_iff -section Image2 - -variable {f : α → β → γ} {s : Set α} {t : Set β} {a : α} {b : β} - -protected theorem Infinite.image2_left (hs : s.Infinite) (hb : b ∈ t) - (hf : InjOn (fun a => f a b) s) : (image2 f s t).Infinite := - (hs.image hf).mono <| image_subset_image2_left hb - -protected theorem Infinite.image2_right (ht : t.Infinite) (ha : a ∈ s) (hf : InjOn (f a) t) : - (image2 f s t).Infinite := - (ht.image hf).mono <| image_subset_image2_right ha - -theorem infinite_image2 (hfs : ∀ b ∈ t, InjOn (fun a => f a b) s) (hft : ∀ a ∈ s, InjOn (f a) t) : - (image2 f s t).Infinite ↔ s.Infinite ∧ t.Nonempty ∨ t.Infinite ∧ s.Nonempty := by - refine ⟨fun h => Set.infinite_prod.1 ?_, ?_⟩ - · rw [← image_uncurry_prod] at h - exact h.of_image _ - · rintro (⟨hs, b, hb⟩ | ⟨ht, a, ha⟩) - · exact hs.image2_left hb (hfs _ hb) - · exact ht.image2_right ha (hft _ ha) - -lemma finite_image2 (hfs : ∀ b ∈ t, InjOn (f · b) s) (hft : ∀ a ∈ s, InjOn (f a) t) : - (image2 f s t).Finite ↔ s.Finite ∧ t.Finite ∨ s = ∅ ∨ t = ∅ := by - rw [← not_infinite, infinite_image2 hfs hft] - simp [not_or, -not_and, not_and_or, not_nonempty_iff_eq_empty] - aesop - -end Image2 - theorem infinite_of_injOn_mapsTo {s : Set α} {t : Set β} {f : α → β} (hi : InjOn f s) (hm : MapsTo f s t) (hs : s.Infinite) : t.Infinite := ((infinite_image_iff hi).2 hs).mono (mapsTo'.mp hm) @@ -1269,22 +860,6 @@ theorem not_injOn_infinite_finite_image {f : α → β} {s : Set α} (h_inf : s. contrapose! h rwa [injective_codRestrict, ← injOn_iff_injective] -theorem infinite_iUnion {ι : Type*} [Infinite ι] {s : ι → Set α} (hs : Function.Injective s) : - (⋃ i, s i).Infinite := - fun hfin ↦ @not_injective_infinite_finite ι _ _ hfin.finite_subsets.to_subtype - (fun i ↦ ⟨s i, subset_iUnion _ _⟩) fun i j h_eq ↦ hs (by simpa using h_eq) - -theorem Infinite.biUnion {ι : Type*} {s : ι → Set α} {a : Set ι} (ha : a.Infinite) - (hs : a.InjOn s) : (⋃ i ∈ a, s i).Infinite := by - rw [biUnion_eq_iUnion] - have _ := ha.to_subtype - exact infinite_iUnion fun ⟨i,hi⟩ ⟨j,hj⟩ hij ↦ by simp [hs hi hj hij] - -theorem Infinite.sUnion {s : Set (Set α)} (hs : s.Infinite) : (⋃₀ s).Infinite := by - rw [sUnion_eq_iUnion] - have _ := hs.to_subtype - exact infinite_iUnion Subtype.coe_injective - /-! ### Order properties -/ section Preorder @@ -1320,137 +895,6 @@ theorem Finite.exists_lt_map_eq_of_forall_mem [LinearOrder α] [Infinite α] {t obtain ⟨a, -, b, -, h⟩ := infinite_univ.exists_lt_map_eq_of_mapsTo hf ht exact ⟨a, b, h⟩ -theorem exists_min_image [LinearOrder β] (s : Set α) (f : α → β) (h1 : s.Finite) : - s.Nonempty → ∃ a ∈ s, ∀ b ∈ s, f a ≤ f b - | ⟨x, hx⟩ => by - simpa only [exists_prop, Finite.mem_toFinset] using - h1.toFinset.exists_min_image f ⟨x, h1.mem_toFinset.2 hx⟩ - -theorem exists_max_image [LinearOrder β] (s : Set α) (f : α → β) (h1 : s.Finite) : - s.Nonempty → ∃ a ∈ s, ∀ b ∈ s, f b ≤ f a - | ⟨x, hx⟩ => by - simpa only [exists_prop, Finite.mem_toFinset] using - h1.toFinset.exists_max_image f ⟨x, h1.mem_toFinset.2 hx⟩ - -theorem exists_lower_bound_image [Nonempty α] [LinearOrder β] (s : Set α) (f : α → β) - (h : s.Finite) : ∃ a : α, ∀ b ∈ s, f a ≤ f b := by - rcases s.eq_empty_or_nonempty with rfl | hs - · exact ‹Nonempty α›.elim fun a => ⟨a, fun _ => False.elim⟩ - · rcases Set.exists_min_image s f h hs with ⟨x₀, _, hx₀⟩ - exact ⟨x₀, fun x hx => hx₀ x hx⟩ - -theorem exists_upper_bound_image [Nonempty α] [LinearOrder β] (s : Set α) (f : α → β) - (h : s.Finite) : ∃ a : α, ∀ b ∈ s, f b ≤ f a := - exists_lower_bound_image (β := βᵒᵈ) s f h - -lemma map_finite_biSup {F ι : Type*} [CompleteLattice α] [CompleteLattice β] [FunLike F α β] - [SupBotHomClass F α β] {s : Set ι} (hs : s.Finite) (f : F) (g : ι → α) : - f (⨆ x ∈ s, g x) = ⨆ x ∈ s, f (g x) := by - have := map_finset_sup f hs.toFinset g - simp only [Finset.sup_eq_iSup, hs.mem_toFinset, comp_apply] at this - exact this - -lemma map_finite_biInf {F ι : Type*} [CompleteLattice α] [CompleteLattice β] [FunLike F α β] - [InfTopHomClass F α β] {s : Set ι} (hs : s.Finite) (f : F) (g : ι → α) : - f (⨅ x ∈ s, g x) = ⨅ x ∈ s, f (g x) := by - have := map_finset_inf f hs.toFinset g - simp only [Finset.inf_eq_iInf, hs.mem_toFinset, comp_apply] at this - exact this - -lemma map_finite_iSup {F ι : Type*} [CompleteLattice α] [CompleteLattice β] [FunLike F α β] - [SupBotHomClass F α β] [Finite ι] (f : F) (g : ι → α) : - f (⨆ i, g i) = ⨆ i, f (g i) := by - rw [← iSup_univ (f := g), ← iSup_univ (f := fun i ↦ f (g i))] - exact map_finite_biSup finite_univ f g - -lemma map_finite_iInf {F ι : Type*} [CompleteLattice α] [CompleteLattice β] [FunLike F α β] - [InfTopHomClass F α β] [Finite ι] (f : F) (g : ι → α) : - f (⨅ i, g i) = ⨅ i, f (g i) := by - rw [← iInf_univ (f := g), ← iInf_univ (f := fun i ↦ f (g i))] - exact map_finite_biInf finite_univ f g - -theorem Finite.iSup_biInf_of_monotone {ι ι' α : Type*} [Preorder ι'] [Nonempty ι'] - [IsDirected ι' (· ≤ ·)] [Order.Frame α] {s : Set ι} (hs : s.Finite) {f : ι → ι' → α} - (hf : ∀ i ∈ s, Monotone (f i)) : ⨆ j, ⨅ i ∈ s, f i j = ⨅ i ∈ s, ⨆ j, f i j := by - induction s, hs using Set.Finite.dinduction_on with - | H0 => simp [iSup_const] - | H1 _ _ ihs => - rw [forall_mem_insert] at hf - simp only [iInf_insert, ← ihs hf.2] - exact iSup_inf_of_monotone hf.1 fun j₁ j₂ hj => iInf₂_mono fun i hi => hf.2 i hi hj - -theorem Finite.iSup_biInf_of_antitone {ι ι' α : Type*} [Preorder ι'] [Nonempty ι'] - [IsDirected ι' (swap (· ≤ ·))] [Order.Frame α] {s : Set ι} (hs : s.Finite) {f : ι → ι' → α} - (hf : ∀ i ∈ s, Antitone (f i)) : ⨆ j, ⨅ i ∈ s, f i j = ⨅ i ∈ s, ⨆ j, f i j := - @Finite.iSup_biInf_of_monotone ι ι'ᵒᵈ α _ _ _ _ _ hs _ fun i hi => (hf i hi).dual_left - -theorem Finite.iInf_biSup_of_monotone {ι ι' α : Type*} [Preorder ι'] [Nonempty ι'] - [IsDirected ι' (swap (· ≤ ·))] [Order.Coframe α] {s : Set ι} (hs : s.Finite) {f : ι → ι' → α} - (hf : ∀ i ∈ s, Monotone (f i)) : ⨅ j, ⨆ i ∈ s, f i j = ⨆ i ∈ s, ⨅ j, f i j := - hs.iSup_biInf_of_antitone (α := αᵒᵈ) fun i hi => (hf i hi).dual_right - -theorem Finite.iInf_biSup_of_antitone {ι ι' α : Type*} [Preorder ι'] [Nonempty ι'] - [IsDirected ι' (· ≤ ·)] [Order.Coframe α] {s : Set ι} (hs : s.Finite) {f : ι → ι' → α} - (hf : ∀ i ∈ s, Antitone (f i)) : ⨅ j, ⨆ i ∈ s, f i j = ⨆ i ∈ s, ⨅ j, f i j := - hs.iSup_biInf_of_monotone (α := αᵒᵈ) fun i hi => (hf i hi).dual_right - -theorem iSup_iInf_of_monotone {ι ι' α : Type*} [Finite ι] [Preorder ι'] [Nonempty ι'] - [IsDirected ι' (· ≤ ·)] [Order.Frame α] {f : ι → ι' → α} (hf : ∀ i, Monotone (f i)) : - ⨆ j, ⨅ i, f i j = ⨅ i, ⨆ j, f i j := by - simpa only [iInf_univ] using finite_univ.iSup_biInf_of_monotone fun i _ => hf i - -theorem iSup_iInf_of_antitone {ι ι' α : Type*} [Finite ι] [Preorder ι'] [Nonempty ι'] - [IsDirected ι' (swap (· ≤ ·))] [Order.Frame α] {f : ι → ι' → α} (hf : ∀ i, Antitone (f i)) : - ⨆ j, ⨅ i, f i j = ⨅ i, ⨆ j, f i j := - @iSup_iInf_of_monotone ι ι'ᵒᵈ α _ _ _ _ _ _ fun i => (hf i).dual_left - -theorem iInf_iSup_of_monotone {ι ι' α : Type*} [Finite ι] [Preorder ι'] [Nonempty ι'] - [IsDirected ι' (swap (· ≤ ·))] [Order.Coframe α] {f : ι → ι' → α} (hf : ∀ i, Monotone (f i)) : - ⨅ j, ⨆ i, f i j = ⨆ i, ⨅ j, f i j := - iSup_iInf_of_antitone (α := αᵒᵈ) fun i => (hf i).dual_right - -theorem iInf_iSup_of_antitone {ι ι' α : Type*} [Finite ι] [Preorder ι'] [Nonempty ι'] - [IsDirected ι' (· ≤ ·)] [Order.Coframe α] {f : ι → ι' → α} (hf : ∀ i, Antitone (f i)) : - ⨅ j, ⨆ i, f i j = ⨆ i, ⨅ j, f i j := - iSup_iInf_of_monotone (α := αᵒᵈ) fun i => (hf i).dual_right - -/-- An increasing union distributes over finite intersection. -/ -theorem iUnion_iInter_of_monotone {ι ι' α : Type*} [Finite ι] [Preorder ι'] [IsDirected ι' (· ≤ ·)] - [Nonempty ι'] {s : ι → ι' → Set α} (hs : ∀ i, Monotone (s i)) : - ⋃ j : ι', ⋂ i : ι, s i j = ⋂ i : ι, ⋃ j : ι', s i j := - iSup_iInf_of_monotone hs - -/-- A decreasing union distributes over finite intersection. -/ -theorem iUnion_iInter_of_antitone {ι ι' α : Type*} [Finite ι] [Preorder ι'] - [IsDirected ι' (swap (· ≤ ·))] [Nonempty ι'] {s : ι → ι' → Set α} (hs : ∀ i, Antitone (s i)) : - ⋃ j : ι', ⋂ i : ι, s i j = ⋂ i : ι, ⋃ j : ι', s i j := - iSup_iInf_of_antitone hs - -/-- An increasing intersection distributes over finite union. -/ -theorem iInter_iUnion_of_monotone {ι ι' α : Type*} [Finite ι] [Preorder ι'] - [IsDirected ι' (swap (· ≤ ·))] [Nonempty ι'] {s : ι → ι' → Set α} (hs : ∀ i, Monotone (s i)) : - ⋂ j : ι', ⋃ i : ι, s i j = ⋃ i : ι, ⋂ j : ι', s i j := - iInf_iSup_of_monotone hs - -/-- A decreasing intersection distributes over finite union. -/ -theorem iInter_iUnion_of_antitone {ι ι' α : Type*} [Finite ι] [Preorder ι'] [IsDirected ι' (· ≤ ·)] - [Nonempty ι'] {s : ι → ι' → Set α} (hs : ∀ i, Antitone (s i)) : - ⋂ j : ι', ⋃ i : ι, s i j = ⋃ i : ι, ⋂ j : ι', s i j := - iInf_iSup_of_antitone hs - -theorem iUnion_pi_of_monotone {ι ι' : Type*} [LinearOrder ι'] [Nonempty ι'] {α : ι → Type*} - {I : Set ι} {s : ∀ i, ι' → Set (α i)} (hI : I.Finite) (hs : ∀ i ∈ I, Monotone (s i)) : - ⋃ j : ι', I.pi (fun i => s i j) = I.pi fun i => ⋃ j, s i j := by - simp only [pi_def, biInter_eq_iInter, preimage_iUnion] - haveI := hI.fintype.finite - refine iUnion_iInter_of_monotone (ι' := ι') (fun (i : I) j₁ j₂ h => ?_) - exact preimage_mono <| hs i i.2 h - -theorem iUnion_univ_pi_of_monotone {ι ι' : Type*} [LinearOrder ι'] [Nonempty ι'] [Finite ι] - {α : ι → Type*} {s : ∀ i, ι' → Set (α i)} (hs : ∀ i, Monotone (s i)) : - ⋃ j : ι', pi univ (fun i => s i j) = pi univ fun i => ⋃ j, s i j := - iUnion_pi_of_monotone finite_univ fun i _ => hs i - theorem finite_range_findGreatest {P : α → ℕ → Prop} [∀ x, DecidablePred (P x)] {b : ℕ} : (range fun x => Nat.findGreatest (P x) b).Finite := (finite_le_nat b).subset <| range_subset_iff.2 fun _ => Nat.findGreatest_le _ @@ -1491,42 +935,6 @@ lemma Finite.exists_minimal_wrt' [PartialOrder β] (f : α → β) (s : Set α) (hs : s.Nonempty) : (∃ a ∈ s, ∀ (a' : α), a' ∈ s → f a' ≤ f a → f a = f a') := Set.Finite.exists_maximal_wrt' (β := βᵒᵈ) f s h hs -section - -variable [Preorder α] [IsDirected α (· ≤ ·)] [Nonempty α] {s : Set α} - -/-- A finite set is bounded above. -/ -protected theorem Finite.bddAbove (hs : s.Finite) : BddAbove s := - Finite.induction_on hs bddAbove_empty fun _ _ h => h.insert _ - -/-- A finite union of sets which are all bounded above is still bounded above. -/ -theorem Finite.bddAbove_biUnion {I : Set β} {S : β → Set α} (H : I.Finite) : - BddAbove (⋃ i ∈ I, S i) ↔ ∀ i ∈ I, BddAbove (S i) := - Finite.induction_on H (by simp only [biUnion_empty, bddAbove_empty, forall_mem_empty]) - fun _ _ hs => by simp only [biUnion_insert, forall_mem_insert, bddAbove_union, hs] - -theorem infinite_of_not_bddAbove : ¬BddAbove s → s.Infinite := - mt Finite.bddAbove - -end - -section - -variable [Preorder α] [IsDirected α (· ≥ ·)] [Nonempty α] {s : Set α} - -/-- A finite set is bounded below. -/ -protected theorem Finite.bddBelow (hs : s.Finite) : BddBelow s := - Finite.bddAbove (α := αᵒᵈ) hs - -/-- A finite union of sets which are all bounded below is still bounded below. -/ -theorem Finite.bddBelow_biUnion {I : Set β} {S : β → Set α} (H : I.Finite) : - BddBelow (⋃ i ∈ I, S i) ↔ ∀ i ∈ I, BddBelow (S i) := - Finite.bddAbove_biUnion (α := αᵒᵈ) H - -theorem infinite_of_not_bddBelow : ¬BddBelow s → s.Infinite := mt Finite.bddBelow - -end - end Set namespace Finset @@ -1539,14 +947,6 @@ lemma exists_card_eq [Infinite α] : ∀ n : ℕ, ∃ s : Finset α, s.card = n obtain ⟨a, ha⟩ := s.exists_not_mem exact ⟨insert a s, card_insert_of_not_mem ha⟩ -/-- A finset is bounded above. -/ -protected theorem bddAbove [SemilatticeSup α] [Nonempty α] (s : Finset α) : BddAbove (↑s : Set α) := - s.finite_toSet.bddAbove - -/-- A finset is bounded below. -/ -protected theorem bddBelow [SemilatticeInf α] [Nonempty α] (s : Finset α) : BddBelow (↑s : Set α) := - s.finite_toSet.bddBelow - end Finset section LinearOrder @@ -1565,13 +965,6 @@ lemma Set.finite_of_forall_not_lt_lt (h : ∀ x ∈ s, ∀ y ∈ s, ∀ z ∈ s, Set.Finite s := @Set.toFinite _ s <| Finite.of_forall_not_lt_lt <| by simpa only [SetCoe.forall'] using h -lemma Set.finite_diff_iUnion_Ioo (s : Set α) : (s \ ⋃ (x ∈ s) (y ∈ s), Ioo x y).Finite := - Set.finite_of_forall_not_lt_lt fun _x hx _y hy _z hz hxy hyz => hy.2 <| mem_iUnion₂_of_mem hx.1 <| - mem_iUnion₂_of_mem hz.1 ⟨hxy, hyz⟩ - -lemma Set.finite_diff_iUnion_Ioo' (s : Set α) : (s \ ⋃ x : s × s, Ioo x.1 x.2).Finite := by - simpa only [iUnion, iSup_prod, iSup_subtype] using s.finite_diff_iUnion_Ioo - lemma Directed.exists_mem_subset_of_finset_subset_biUnion {α ι : Type*} [Nonempty ι] {f : ι → Set α} (h : Directed (· ⊆ ·) f) {s : Finset α} (hs : (s : Set α) ⊆ ⋃ i, f i) : ∃ i, (s : Set α) ⊆ f i := by @@ -1583,26 +976,4 @@ lemma Directed.exists_mem_subset_of_finset_subset_biUnion {α ι : Type*} [Nonem rcases h i j with ⟨k, hik, hjk⟩ exact ⟨k, hik hi, hj.trans hjk⟩ -theorem DirectedOn.exists_mem_subset_of_finset_subset_biUnion {α ι : Type*} {f : ι → Set α} - {c : Set ι} (hn : c.Nonempty) (hc : DirectedOn (fun i j => f i ⊆ f j) c) {s : Finset α} - (hs : (s : Set α) ⊆ ⋃ i ∈ c, f i) : ∃ i ∈ c, (s : Set α) ⊆ f i := by - rw [Set.biUnion_eq_iUnion] at hs - 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 := Mathlib.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 diff --git a/Mathlib/Data/Set/Finite/Lattice.lean b/Mathlib/Data/Set/Finite/Lattice.lean new file mode 100644 index 0000000000000..d48b9b82b810e --- /dev/null +++ b/Mathlib/Data/Set/Finite/Lattice.lean @@ -0,0 +1,399 @@ +/- +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.Finite.Sigma +import Mathlib.Data.Set.Finite.Powerset +import Mathlib.Data.Set.Finite.Range + +/-! +# Finiteness of unions and intersections + +## Implementation notes + +Each result in this file should come in three forms: a `Fintype` instance, a `Finite` instance +and a `Set.Finite` constructor. + +## Tags + +finite sets +-/ + +assert_not_exists OrderedRing +assert_not_exists MonoidWithZero + +open Set Function + +universe u v w x + +variable {α : Type u} {β : Type v} {ι : Sort w} {γ : Type x} + +namespace Set + +/-! ### Fintype instances + +Every instance here should have a corresponding `Set.Finite` constructor in the next section. +-/ + +section FintypeInstances + +instance fintypeiUnion [DecidableEq α] [Fintype (PLift ι)] (f : ι → Set α) [∀ i, Fintype (f i)] : + Fintype (⋃ i, f i) := + Fintype.ofFinset (Finset.univ.biUnion fun i : PLift ι => (f i.down).toFinset) <| by simp + +instance fintypesUnion [DecidableEq α] {s : Set (Set α)} [Fintype s] + [H : ∀ t : s, Fintype (t : Set α)] : Fintype (⋃₀ s) := by + rw [sUnion_eq_iUnion] + exact @Set.fintypeiUnion _ _ _ _ _ H + +lemma toFinset_iUnion [Fintype β] [DecidableEq α] (f : β → Set α) + [∀ w, Fintype (f w)] : + Set.toFinset (⋃ (x : β), f x) = + Finset.biUnion (Finset.univ : Finset β) (fun x => (f x).toFinset) := by + ext v + simp only [mem_toFinset, mem_iUnion, Finset.mem_biUnion, Finset.mem_univ, true_and] + +end FintypeInstances + +end Set + +/-! ### Finite instances + +There is seemingly some overlap between the following instances and the `Fintype` instances +in `Data.Set.Finite`. While every `Fintype` instance gives a `Finite` instance, those +instances that depend on `Fintype` or `Decidable` instances need an additional `Finite` instance +to be able to generally apply. + +Some set instances do not appear here since they are consequences of others, for example +`Subtype.Finite` for subsets of a finite type. +-/ + + +namespace Finite.Set + +open scoped Classical + +instance finite_iUnion [Finite ι] (f : ι → Set α) [∀ i, Finite (f i)] : Finite (⋃ i, f i) := by + rw [iUnion_eq_range_psigma] + apply Set.finite_range + +instance finite_sUnion {s : Set (Set α)} [Finite s] [H : ∀ t : s, Finite (t : Set α)] : + Finite (⋃₀ s) := by + rw [sUnion_eq_iUnion] + exact @Finite.Set.finite_iUnion _ _ _ _ H + +theorem finite_biUnion {ι : Type*} (s : Set ι) [Finite s] (t : ι → Set α) + (H : ∀ i ∈ s, Finite (t i)) : Finite (⋃ x ∈ s, t x) := by + rw [biUnion_eq_iUnion] + haveI : ∀ i : s, Finite (t i) := fun i => H i i.property + infer_instance + +instance finite_biUnion' {ι : Type*} (s : Set ι) [Finite s] (t : ι → Set α) [∀ i, Finite (t i)] : + Finite (⋃ x ∈ s, t x) := + finite_biUnion s t fun _ _ => inferInstance + +/-- Example: `Finite (⋃ (i < n), f i)` where `f : ℕ → Set α` and `[∀ i, Finite (f i)]` +(when given instances from `Order.Interval.Finset.Nat`). +-/ +instance finite_biUnion'' {ι : Type*} (p : ι → Prop) [h : Finite { x | p x }] (t : ι → Set α) + [∀ i, Finite (t i)] : Finite (⋃ (x) (_ : p x), t x) := + @Finite.Set.finite_biUnion' _ _ (setOf p) h t _ + +instance finite_iInter {ι : Sort*} [Nonempty ι] (t : ι → Set α) [∀ i, Finite (t i)] : + Finite (⋂ i, t i) := + Finite.Set.subset (t <| Classical.arbitrary ι) (iInter_subset _ _) + +end Finite.Set + +namespace Set + +/-! ### Constructors for `Set.Finite` + +Every constructor here should have a corresponding `Fintype` instance in the previous section +(or in the `Fintype` module). + +The implementation of these constructors ideally should be no more than `Set.toFinite`, +after possibly setting up some `Fintype` and classical `Decidable` instances. +-/ + + +section SetFiniteConstructors + +theorem finite_iUnion [Finite ι] {f : ι → Set α} (H : ∀ i, (f i).Finite) : (⋃ i, f i).Finite := + haveI := fun i => (H i).to_subtype + toFinite _ + +/-- Dependent version of `Finite.biUnion`. -/ +theorem Finite.biUnion' {ι} {s : Set ι} (hs : s.Finite) {t : ∀ i ∈ s, Set α} + (ht : ∀ i (hi : i ∈ s), (t i hi).Finite) : (⋃ i ∈ s, t i ‹_›).Finite := by + have := hs.to_subtype + rw [biUnion_eq_iUnion] + apply finite_iUnion fun i : s => ht i.1 i.2 + +theorem Finite.biUnion {ι} {s : Set ι} (hs : s.Finite) {t : ι → Set α} + (ht : ∀ i ∈ s, (t i).Finite) : (⋃ i ∈ s, t i).Finite := + hs.biUnion' ht + +theorem Finite.sUnion {s : Set (Set α)} (hs : s.Finite) (H : ∀ t ∈ s, Set.Finite t) : + (⋃₀ s).Finite := by + simpa only [sUnion_eq_biUnion] using hs.biUnion H + +theorem Finite.sInter {α : Type*} {s : Set (Set α)} {t : Set α} (ht : t ∈ s) (hf : t.Finite) : + (⋂₀ s).Finite := + hf.subset (sInter_subset_of_mem ht) + +/-- If sets `s i` are finite for all `i` from a finite set `t` and are empty for `i ∉ t`, then the +union `⋃ i, s i` is a finite set. -/ +theorem Finite.iUnion {ι : Type*} {s : ι → Set α} {t : Set ι} (ht : t.Finite) + (hs : ∀ i ∈ t, (s i).Finite) (he : ∀ i, i ∉ t → s i = ∅) : (⋃ i, s i).Finite := by + suffices ⋃ i, s i ⊆ ⋃ i ∈ t, s i by exact (ht.biUnion hs).subset this + refine iUnion_subset fun i x hx => ?_ + by_cases hi : i ∈ t + · exact mem_biUnion hi hx + · rw [he i hi, mem_empty_iff_false] at hx + contradiction + +section preimage +variable {f : α → β} {s : Set β} + +theorem Finite.preimage' (h : s.Finite) (hf : ∀ b ∈ s, (f ⁻¹' {b}).Finite) : + (f ⁻¹' s).Finite := by + rw [← Set.biUnion_preimage_singleton] + exact Set.Finite.biUnion h hf + +end preimage + +/-- A finite union of finsets is finite. -/ +theorem union_finset_finite_of_range_finite (f : α → Finset β) (h : (range f).Finite) : + (⋃ a, (f a : Set β)).Finite := by + rw [← biUnion_range] + exact h.biUnion fun y _ => y.finite_toSet + +end SetFiniteConstructors + +/-! ### Properties -/ + +theorem finite_subset_iUnion {s : Set α} (hs : s.Finite) {ι} {t : ι → Set α} (h : s ⊆ ⋃ i, t i) : + ∃ I : Set ι, I.Finite ∧ s ⊆ ⋃ i ∈ I, t i := by + have := hs.to_subtype + choose f hf using show ∀ x : s, ∃ i, x.1 ∈ t i by simpa [subset_def] using h + refine ⟨range f, finite_range f, fun x hx => ?_⟩ + rw [biUnion_range, mem_iUnion] + exact ⟨⟨x, hx⟩, hf _⟩ + +theorem eq_finite_iUnion_of_finite_subset_iUnion {ι} {s : ι → Set α} {t : Set α} (tfin : t.Finite) + (h : t ⊆ ⋃ i, s i) : + ∃ I : Set ι, + I.Finite ∧ + ∃ σ : { i | i ∈ I } → Set α, (∀ i, (σ i).Finite) ∧ (∀ i, σ i ⊆ s i) ∧ t = ⋃ i, σ i := + let ⟨I, Ifin, hI⟩ := finite_subset_iUnion tfin h + ⟨I, Ifin, fun x => s x ∩ t, fun _ => tfin.subset inter_subset_right, fun _ => + inter_subset_left, by + ext x + rw [mem_iUnion] + constructor + · intro x_in + rcases mem_iUnion.mp (hI x_in) with ⟨i, _, ⟨hi, rfl⟩, H⟩ + exact ⟨⟨i, hi⟩, ⟨H, x_in⟩⟩ + · rintro ⟨i, -, H⟩ + exact H⟩ + +/-! ### Infinite sets -/ + +variable {s t : Set α} + +theorem infinite_iUnion {ι : Type*} [Infinite ι] {s : ι → Set α} (hs : Function.Injective s) : + (⋃ i, s i).Infinite := + fun hfin ↦ @not_injective_infinite_finite ι _ _ hfin.finite_subsets.to_subtype + (fun i ↦ ⟨s i, subset_iUnion _ _⟩) fun i j h_eq ↦ hs (by simpa using h_eq) + +theorem Infinite.biUnion {ι : Type*} {s : ι → Set α} {a : Set ι} (ha : a.Infinite) + (hs : a.InjOn s) : (⋃ i ∈ a, s i).Infinite := by + rw [biUnion_eq_iUnion] + have _ := ha.to_subtype + exact infinite_iUnion fun ⟨i,hi⟩ ⟨j,hj⟩ hij ↦ by simp [hs hi hj hij] + +theorem Infinite.sUnion {s : Set (Set α)} (hs : s.Infinite) : (⋃₀ s).Infinite := by + rw [sUnion_eq_iUnion] + have _ := hs.to_subtype + exact infinite_iUnion Subtype.coe_injective + +/-! ### Order properties -/ + +lemma map_finite_biSup {F ι : Type*} [CompleteLattice α] [CompleteLattice β] [FunLike F α β] + [SupBotHomClass F α β] {s : Set ι} (hs : s.Finite) (f : F) (g : ι → α) : + f (⨆ x ∈ s, g x) = ⨆ x ∈ s, f (g x) := by + have := map_finset_sup f hs.toFinset g + simp only [Finset.sup_eq_iSup, hs.mem_toFinset, comp_apply] at this + exact this + +lemma map_finite_biInf {F ι : Type*} [CompleteLattice α] [CompleteLattice β] [FunLike F α β] + [InfTopHomClass F α β] {s : Set ι} (hs : s.Finite) (f : F) (g : ι → α) : + f (⨅ x ∈ s, g x) = ⨅ x ∈ s, f (g x) := by + have := map_finset_inf f hs.toFinset g + simp only [Finset.inf_eq_iInf, hs.mem_toFinset, comp_apply] at this + exact this + +lemma map_finite_iSup {F ι : Type*} [CompleteLattice α] [CompleteLattice β] [FunLike F α β] + [SupBotHomClass F α β] [Finite ι] (f : F) (g : ι → α) : + f (⨆ i, g i) = ⨆ i, f (g i) := by + rw [← iSup_univ (f := g), ← iSup_univ (f := fun i ↦ f (g i))] + exact map_finite_biSup finite_univ f g + +lemma map_finite_iInf {F ι : Type*} [CompleteLattice α] [CompleteLattice β] [FunLike F α β] + [InfTopHomClass F α β] [Finite ι] (f : F) (g : ι → α) : + f (⨅ i, g i) = ⨅ i, f (g i) := by + rw [← iInf_univ (f := g), ← iInf_univ (f := fun i ↦ f (g i))] + exact map_finite_biInf finite_univ f g + +theorem Finite.iSup_biInf_of_monotone {ι ι' α : Type*} [Preorder ι'] [Nonempty ι'] + [IsDirected ι' (· ≤ ·)] [Order.Frame α] {s : Set ι} (hs : s.Finite) {f : ι → ι' → α} + (hf : ∀ i ∈ s, Monotone (f i)) : ⨆ j, ⨅ i ∈ s, f i j = ⨅ i ∈ s, ⨆ j, f i j := by + induction s, hs using Set.Finite.dinduction_on with + | H0 => simp [iSup_const] + | H1 _ _ ihs => + rw [forall_mem_insert] at hf + simp only [iInf_insert, ← ihs hf.2] + exact iSup_inf_of_monotone hf.1 fun j₁ j₂ hj => iInf₂_mono fun i hi => hf.2 i hi hj + +theorem Finite.iSup_biInf_of_antitone {ι ι' α : Type*} [Preorder ι'] [Nonempty ι'] + [IsDirected ι' (swap (· ≤ ·))] [Order.Frame α] {s : Set ι} (hs : s.Finite) {f : ι → ι' → α} + (hf : ∀ i ∈ s, Antitone (f i)) : ⨆ j, ⨅ i ∈ s, f i j = ⨅ i ∈ s, ⨆ j, f i j := + @Finite.iSup_biInf_of_monotone ι ι'ᵒᵈ α _ _ _ _ _ hs _ fun i hi => (hf i hi).dual_left + +theorem Finite.iInf_biSup_of_monotone {ι ι' α : Type*} [Preorder ι'] [Nonempty ι'] + [IsDirected ι' (swap (· ≤ ·))] [Order.Coframe α] {s : Set ι} (hs : s.Finite) {f : ι → ι' → α} + (hf : ∀ i ∈ s, Monotone (f i)) : ⨅ j, ⨆ i ∈ s, f i j = ⨆ i ∈ s, ⨅ j, f i j := + hs.iSup_biInf_of_antitone (α := αᵒᵈ) fun i hi => (hf i hi).dual_right + +theorem Finite.iInf_biSup_of_antitone {ι ι' α : Type*} [Preorder ι'] [Nonempty ι'] + [IsDirected ι' (· ≤ ·)] [Order.Coframe α] {s : Set ι} (hs : s.Finite) {f : ι → ι' → α} + (hf : ∀ i ∈ s, Antitone (f i)) : ⨅ j, ⨆ i ∈ s, f i j = ⨆ i ∈ s, ⨅ j, f i j := + hs.iSup_biInf_of_monotone (α := αᵒᵈ) fun i hi => (hf i hi).dual_right + +theorem iSup_iInf_of_monotone {ι ι' α : Type*} [Finite ι] [Preorder ι'] [Nonempty ι'] + [IsDirected ι' (· ≤ ·)] [Order.Frame α] {f : ι → ι' → α} (hf : ∀ i, Monotone (f i)) : + ⨆ j, ⨅ i, f i j = ⨅ i, ⨆ j, f i j := by + simpa only [iInf_univ] using finite_univ.iSup_biInf_of_monotone fun i _ => hf i + +theorem iSup_iInf_of_antitone {ι ι' α : Type*} [Finite ι] [Preorder ι'] [Nonempty ι'] + [IsDirected ι' (swap (· ≤ ·))] [Order.Frame α] {f : ι → ι' → α} (hf : ∀ i, Antitone (f i)) : + ⨆ j, ⨅ i, f i j = ⨅ i, ⨆ j, f i j := + @iSup_iInf_of_monotone ι ι'ᵒᵈ α _ _ _ _ _ _ fun i => (hf i).dual_left + +theorem iInf_iSup_of_monotone {ι ι' α : Type*} [Finite ι] [Preorder ι'] [Nonempty ι'] + [IsDirected ι' (swap (· ≤ ·))] [Order.Coframe α] {f : ι → ι' → α} (hf : ∀ i, Monotone (f i)) : + ⨅ j, ⨆ i, f i j = ⨆ i, ⨅ j, f i j := + iSup_iInf_of_antitone (α := αᵒᵈ) fun i => (hf i).dual_right + +theorem iInf_iSup_of_antitone {ι ι' α : Type*} [Finite ι] [Preorder ι'] [Nonempty ι'] + [IsDirected ι' (· ≤ ·)] [Order.Coframe α] {f : ι → ι' → α} (hf : ∀ i, Antitone (f i)) : + ⨅ j, ⨆ i, f i j = ⨆ i, ⨅ j, f i j := + iSup_iInf_of_monotone (α := αᵒᵈ) fun i => (hf i).dual_right + +/-- An increasing union distributes over finite intersection. -/ +theorem iUnion_iInter_of_monotone {ι ι' α : Type*} [Finite ι] [Preorder ι'] [IsDirected ι' (· ≤ ·)] + [Nonempty ι'] {s : ι → ι' → Set α} (hs : ∀ i, Monotone (s i)) : + ⋃ j : ι', ⋂ i : ι, s i j = ⋂ i : ι, ⋃ j : ι', s i j := + iSup_iInf_of_monotone hs + +/-- A decreasing union distributes over finite intersection. -/ +theorem iUnion_iInter_of_antitone {ι ι' α : Type*} [Finite ι] [Preorder ι'] + [IsDirected ι' (swap (· ≤ ·))] [Nonempty ι'] {s : ι → ι' → Set α} (hs : ∀ i, Antitone (s i)) : + ⋃ j : ι', ⋂ i : ι, s i j = ⋂ i : ι, ⋃ j : ι', s i j := + iSup_iInf_of_antitone hs + +/-- An increasing intersection distributes over finite union. -/ +theorem iInter_iUnion_of_monotone {ι ι' α : Type*} [Finite ι] [Preorder ι'] + [IsDirected ι' (swap (· ≤ ·))] [Nonempty ι'] {s : ι → ι' → Set α} (hs : ∀ i, Monotone (s i)) : + ⋂ j : ι', ⋃ i : ι, s i j = ⋃ i : ι, ⋂ j : ι', s i j := + iInf_iSup_of_monotone hs + +/-- A decreasing intersection distributes over finite union. -/ +theorem iInter_iUnion_of_antitone {ι ι' α : Type*} [Finite ι] [Preorder ι'] [IsDirected ι' (· ≤ ·)] + [Nonempty ι'] {s : ι → ι' → Set α} (hs : ∀ i, Antitone (s i)) : + ⋂ j : ι', ⋃ i : ι, s i j = ⋃ i : ι, ⋂ j : ι', s i j := + iInf_iSup_of_antitone hs + +theorem iUnion_pi_of_monotone {ι ι' : Type*} [LinearOrder ι'] [Nonempty ι'] {α : ι → Type*} + {I : Set ι} {s : ∀ i, ι' → Set (α i)} (hI : I.Finite) (hs : ∀ i ∈ I, Monotone (s i)) : + ⋃ j : ι', I.pi (fun i => s i j) = I.pi fun i => ⋃ j, s i j := by + simp only [pi_def, biInter_eq_iInter, preimage_iUnion] + haveI := hI.fintype.finite + refine iUnion_iInter_of_monotone (ι' := ι') (fun (i : I) j₁ j₂ h => ?_) + exact preimage_mono <| hs i i.2 h + +theorem iUnion_univ_pi_of_monotone {ι ι' : Type*} [LinearOrder ι'] [Nonempty ι'] [Finite ι] + {α : ι → Type*} {s : ∀ i, ι' → Set (α i)} (hs : ∀ i, Monotone (s i)) : + ⋃ j : ι', pi univ (fun i => s i j) = pi univ fun i => ⋃ j, s i j := + iUnion_pi_of_monotone finite_univ fun i _ => hs i + +section + +variable [Preorder α] [IsDirected α (· ≤ ·)] [Nonempty α] {s : Set α} + +/-- A finite set is bounded above. -/ +protected theorem Finite.bddAbove (hs : s.Finite) : BddAbove s := + Finite.induction_on hs bddAbove_empty fun _ _ h => h.insert _ + +/-- A finite union of sets which are all bounded above is still bounded above. -/ +theorem Finite.bddAbove_biUnion {I : Set β} {S : β → Set α} (H : I.Finite) : + BddAbove (⋃ i ∈ I, S i) ↔ ∀ i ∈ I, BddAbove (S i) := + Finite.induction_on H (by simp only [biUnion_empty, bddAbove_empty, forall_mem_empty]) + fun _ _ hs => by simp only [biUnion_insert, forall_mem_insert, bddAbove_union, hs] + +theorem infinite_of_not_bddAbove : ¬BddAbove s → s.Infinite := + mt Finite.bddAbove + +end + +section + +variable [Preorder α] [IsDirected α (· ≥ ·)] [Nonempty α] {s : Set α} + +/-- A finite set is bounded below. -/ +protected theorem Finite.bddBelow (hs : s.Finite) : BddBelow s := + Finite.bddAbove (α := αᵒᵈ) hs + +/-- A finite union of sets which are all bounded below is still bounded below. -/ +theorem Finite.bddBelow_biUnion {I : Set β} {S : β → Set α} (H : I.Finite) : + BddBelow (⋃ i ∈ I, S i) ↔ ∀ i ∈ I, BddBelow (S i) := + Finite.bddAbove_biUnion (α := αᵒᵈ) H + +theorem infinite_of_not_bddBelow : ¬BddBelow s → s.Infinite := mt Finite.bddBelow + +end + +end Set + +namespace Finset + +/-- A finset is bounded above. -/ +protected theorem bddAbove [SemilatticeSup α] [Nonempty α] (s : Finset α) : BddAbove (↑s : Set α) := + s.finite_toSet.bddAbove + +/-- A finset is bounded below. -/ +protected theorem bddBelow [SemilatticeInf α] [Nonempty α] (s : Finset α) : BddBelow (↑s : Set α) := + s.finite_toSet.bddBelow + +end Finset + +section LinearOrder +variable [LinearOrder α] {s : Set α} + +lemma Set.finite_diff_iUnion_Ioo (s : Set α) : (s \ ⋃ (x ∈ s) (y ∈ s), Ioo x y).Finite := + Set.finite_of_forall_not_lt_lt fun _x hx _y hy _z hz hxy hyz => hy.2 <| mem_iUnion₂_of_mem hx.1 <| + mem_iUnion₂_of_mem hz.1 ⟨hxy, hyz⟩ + +lemma Set.finite_diff_iUnion_Ioo' (s : Set α) : (s \ ⋃ x : s × s, Ioo x.1 x.2).Finite := by + simpa only [iUnion, iSup_prod, iSup_subtype] using s.finite_diff_iUnion_Ioo + +theorem DirectedOn.exists_mem_subset_of_finset_subset_biUnion {α ι : Type*} {f : ι → Set α} + {c : Set ι} (hn : c.Nonempty) (hc : DirectedOn (fun i j => f i ⊆ f j) c) {s : Finset α} + (hs : (s : Set α) ⊆ ⋃ i ∈ c, f i) : ∃ i ∈ c, (s : Set α) ⊆ f i := by + rw [Set.biUnion_eq_iUnion] at hs + haveI := hn.coe_sort + simpa using (directed_comp.2 hc.directed_val).exists_mem_subset_of_finset_subset_biUnion hs + +end LinearOrder diff --git a/Mathlib/Data/Set/Finite/Lemmas.lean b/Mathlib/Data/Set/Finite/Lemmas.lean new file mode 100644 index 0000000000000..29701fca18502 --- /dev/null +++ b/Mathlib/Data/Set/Finite/Lemmas.lean @@ -0,0 +1,92 @@ +/- +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.Max +import Mathlib.Data.Set.Finite.Basic +import Mathlib.Data.Set.Lattice +import Mathlib.Data.Finite.Powerset +import Mathlib.Logic.Embedding.Set + +/-! +# Lemmas on finiteness of sets + +This file should contain lemmas that prove some result under the *assumption* of `Set.Finite`. +If your proof has as *result* `Set.Finite`, then it should go to a more specific file. + +## Tags + +finite sets +-/ + +assert_not_exists OrderedRing +assert_not_exists MonoidWithZero + +open Set Function + +universe u v w x + +variable {α : Type u} {β : Type v} {ι : Sort w} {γ : Type x} + +namespace Set + +/-! ### Properties -/ + +theorem Finite.fin_embedding {s : Set α} (h : s.Finite) : + ∃ (n : ℕ) (f : Fin n ↪ α), range f = s := + ⟨_, (Fintype.equivFin (h.toFinset : Set α)).symm.asEmbedding, by + simp only [Finset.coe_sort_coe, Equiv.asEmbedding_range, Finite.coe_toFinset, setOf_mem_eq]⟩ + +theorem Finite.fin_param {s : Set α} (h : s.Finite) : + ∃ (n : ℕ) (f : Fin n → α), Injective f ∧ range f = s := + let ⟨n, f, hf⟩ := h.fin_embedding + ⟨n, f, f.injective, hf⟩ + +/-- Induction up to a finite set `S`. -/ +theorem Finite.induction_to {C : Set α → Prop} {S : Set α} (h : S.Finite) + (S0 : Set α) (hS0 : S0 ⊆ S) (H0 : C S0) (H1 : ∀ s ⊂ S, C s → ∃ a ∈ S \ s, C (insert a s)) : + C S := by + have : Finite S := Finite.to_subtype h + have : Finite {T : Set α // T ⊆ S} := Finite.of_equiv (Set S) (Equiv.Set.powerset S).symm + rw [← Subtype.coe_mk (p := (· ⊆ S)) _ le_rfl] + rw [← Subtype.coe_mk (p := (· ⊆ S)) _ hS0] at H0 + refine Finite.to_wellFoundedGT.wf.induction_bot' (fun s hs hs' ↦ ?_) H0 + obtain ⟨a, ⟨ha1, ha2⟩, ha'⟩ := H1 s (ssubset_of_ne_of_subset hs s.2) hs' + exact ⟨⟨insert a s.1, insert_subset ha1 s.2⟩, Set.ssubset_insert ha2, ha'⟩ + +/-- Induction up to `univ`. -/ +theorem Finite.induction_to_univ [Finite α] {C : Set α → Prop} (S0 : Set α) + (H0 : C S0) (H1 : ∀ S ≠ univ, C S → ∃ a ∉ S, C (insert a S)) : C univ := + finite_univ.induction_to S0 (subset_univ S0) H0 (by simpa [ssubset_univ_iff]) + +/-! ### Infinite sets -/ + +variable {s t : Set α} + +/-! ### Order properties -/ + +theorem exists_min_image [LinearOrder β] (s : Set α) (f : α → β) (h1 : s.Finite) : + s.Nonempty → ∃ a ∈ s, ∀ b ∈ s, f a ≤ f b + | ⟨x, hx⟩ => by + simpa only [exists_prop, Finite.mem_toFinset] using + h1.toFinset.exists_min_image f ⟨x, h1.mem_toFinset.2 hx⟩ + +theorem exists_max_image [LinearOrder β] (s : Set α) (f : α → β) (h1 : s.Finite) : + s.Nonempty → ∃ a ∈ s, ∀ b ∈ s, f b ≤ f a + | ⟨x, hx⟩ => by + simpa only [exists_prop, Finite.mem_toFinset] using + h1.toFinset.exists_max_image f ⟨x, h1.mem_toFinset.2 hx⟩ + +theorem exists_lower_bound_image [Nonempty α] [LinearOrder β] (s : Set α) (f : α → β) + (h : s.Finite) : ∃ a : α, ∀ b ∈ s, f a ≤ f b := by + rcases s.eq_empty_or_nonempty with rfl | hs + · exact ‹Nonempty α›.elim fun a => ⟨a, fun _ => False.elim⟩ + · rcases Set.exists_min_image s f h hs with ⟨x₀, _, hx₀⟩ + exact ⟨x₀, fun x hx => hx₀ x hx⟩ + +theorem exists_upper_bound_image [Nonempty α] [LinearOrder β] (s : Set α) (f : α → β) + (h : s.Finite) : ∃ a : α, ∀ b ∈ s, f b ≤ f a := + exists_lower_bound_image (β := βᵒᵈ) s f h + +end Set diff --git a/Mathlib/Data/Set/Finite/List.lean b/Mathlib/Data/Set/Finite/List.lean new file mode 100644 index 0000000000000..1282e6b3c5da0 --- /dev/null +++ b/Mathlib/Data/Set/Finite/List.lean @@ -0,0 +1,35 @@ +/- +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.Set.Finite.Basic +import Mathlib.Data.Set.Finite.Lattice +import Mathlib.Data.Set.Finite.Range +import Mathlib.Data.Set.Lattice +import Mathlib.Data.Finite.Sigma +import Mathlib.Data.Finite.Vector + +/-! +# Finiteness of sets of lists + +## Tags + +finite sets +-/ + +assert_not_exists OrderedRing +assert_not_exists MonoidWithZero + +namespace List +variable (α : Type*) [Finite α] (n : ℕ) + +lemma finite_length_eq : {l : List α | l.length = n}.Finite := Mathlib.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 diff --git a/Mathlib/Data/Set/Finite/Monad.lean b/Mathlib/Data/Set/Finite/Monad.lean new file mode 100644 index 0000000000000..4614681a6df64 --- /dev/null +++ b/Mathlib/Data/Set/Finite/Monad.lean @@ -0,0 +1,125 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Mario Carneiro, Kyle Miller +-/ +import Mathlib.Data.Finite.Prod +import Mathlib.Data.Set.Finite.Lattice +import Mathlib.Data.Set.Functor + +/-! +# Finiteness of the Set monad operations + +## Tags + +finite sets +-/ + +assert_not_exists OrderedRing +assert_not_exists MonoidWithZero + +open Set Function + +universe u v w x + +variable {α : Type u} {β : Type v} {ι : Sort w} {γ : Type x} + +namespace Set + +/-! ### Fintype instances + +Every instance here should have a corresponding `Set.Finite` constructor in the next section. +-/ + +section FintypeInstances + +section monad +attribute [local instance] Set.monad + +/-- If `s : Set α` is a set with `Fintype` instance and `f : α → Set β` is a function such that +each `f a`, `a ∈ s`, has a `Fintype` structure, then `s >>= f` has a `Fintype` structure. -/ +def fintypeBind {α β} [DecidableEq β] (s : Set α) [Fintype s] (f : α → Set β) + (H : ∀ a ∈ s, Fintype (f a)) : Fintype (s >>= f) := + Set.fintypeBiUnion s f H + +instance fintypeBind' {α β} [DecidableEq β] (s : Set α) [Fintype s] (f : α → Set β) + [∀ a, Fintype (f a)] : Fintype (s >>= f) := + Set.fintypeBiUnion' s f + +end monad + +instance fintypePure : ∀ a : α, Fintype (pure a : Set α) := + Set.fintypeSingleton + +instance fintypeSeq [DecidableEq β] (f : Set (α → β)) (s : Set α) [Fintype f] [Fintype s] : + Fintype (f.seq s) := by + rw [seq_def] + apply Set.fintypeBiUnion' + +instance fintypeSeq' {α β : Type u} [DecidableEq β] (f : Set (α → β)) (s : Set α) [Fintype f] + [Fintype s] : Fintype (f <*> s) := + Set.fintypeSeq f s + +end FintypeInstances + +end Set + +/-! ### Finite instances + +There is seemingly some overlap between the following instances and the `Fintype` instances +in `Data.Set.Finite`. While every `Fintype` instance gives a `Finite` instance, those +instances that depend on `Fintype` or `Decidable` instances need an additional `Finite` instance +to be able to generally apply. + +Some set instances do not appear here since they are consequences of others, for example +`Subtype.Finite` for subsets of a finite type. +-/ + + +namespace Finite.Set + +open scoped Classical + +theorem finite_pure (a : α) : (pure a : Set α).Finite := + toFinite _ + +instance finite_seq (f : Set (α → β)) (s : Set α) [Finite f] [Finite s] : Finite (f.seq s) := by + rw [seq_def] + infer_instance + +end Finite.Set + +namespace Set + +/-! ### Constructors for `Set.Finite` + +Every constructor here should have a corresponding `Fintype` instance in the previous section +(or in the `Fintype` module). + +The implementation of these constructors ideally should be no more than `Set.toFinite`, +after possibly setting up some `Fintype` and classical `Decidable` instances. +-/ + + +section SetFiniteConstructors + +section monad +attribute [local instance] Set.monad + +theorem Finite.bind {α β} {s : Set α} {f : α → Set β} (h : s.Finite) (hf : ∀ a ∈ s, (f a).Finite) : + (s >>= f).Finite := + h.biUnion hf + +end monad + +theorem Finite.seq {f : Set (α → β)} {s : Set α} (hf : f.Finite) (hs : s.Finite) : + (f.seq s).Finite := + hf.image2 _ hs + +theorem Finite.seq' {α β : Type u} {f : Set (α → β)} {s : Set α} (hf : f.Finite) (hs : s.Finite) : + (f <*> s).Finite := + hf.seq hs + +end SetFiniteConstructors + +end Set diff --git a/Mathlib/Data/Set/Finite/Powerset.lean b/Mathlib/Data/Set/Finite/Powerset.lean new file mode 100644 index 0000000000000..55f30e0cb51b6 --- /dev/null +++ b/Mathlib/Data/Set/Finite/Powerset.lean @@ -0,0 +1,56 @@ +/- +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.Powerset +import Mathlib.Data.Set.Finite.Basic + +/-! +# Finiteness of the powerset of a finite set + +## Implementation notes + +Each result in this file should come in three forms: a `Fintype` instance, a `Finite` instance +and a `Set.Finite` constructor. + +## Tags + +finite sets +-/ + +assert_not_exists OrderedRing +assert_not_exists MonoidWithZero + +open Set Function + +universe u v w x + +variable {α : Type u} {β : Type v} {ι : Sort w} {γ : Type x} + +namespace Set + +/-! ### Constructors for `Set.Finite` + +Every constructor here should have a corresponding `Fintype` instance in the `Fintype` module. + +The implementation of these constructors ideally should be no more than `Set.toFinite`, +after possibly setting up some `Fintype` and classical `Decidable` instances. +-/ + + +section SetFiniteConstructors + +/-- There are finitely many subsets of a given finite set -/ +theorem Finite.finite_subsets {α : Type u} {a : Set α} (h : a.Finite) : { b | b ⊆ a }.Finite := by + convert ((Finset.powerset h.toFinset).map Finset.coeEmb.1).finite_toSet + ext s + simpa [← @exists_finite_iff_finset α fun t => t ⊆ a ∧ t = s, Finite.subset_toFinset, + ← and_assoc, Finset.coeEmb] using h.subset + +protected theorem Finite.powerset {s : Set α} (h : s.Finite) : (𝒫 s).Finite := + h.finite_subsets + +end SetFiniteConstructors + +end Set diff --git a/Mathlib/Data/Set/Finite/Range.lean b/Mathlib/Data/Set/Finite/Range.lean new file mode 100644 index 0000000000000..fdbe1c806503c --- /dev/null +++ b/Mathlib/Data/Set/Finite/Range.lean @@ -0,0 +1,97 @@ +/- +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.Fintype.Card +import Mathlib.Data.ULift + +/-! +# Finiteness of `Set.range` + +## Implementation notes + +Each result in this file should come in three forms: a `Fintype` instance, a `Finite` instance +and a `Set.Finite` constructor. + +## Tags + +finite sets +-/ + +assert_not_exists OrderedRing +assert_not_exists MonoidWithZero + +open Set Function + +universe u v w x + +variable {α : Type u} {β : Type v} {ι : Sort w} {γ : Type x} + +namespace Set + +/-! ### Fintype instances + +Every instance here should have a corresponding `Set.Finite` constructor in the next section. +-/ + +section FintypeInstances + +instance fintypeRange [DecidableEq α] (f : ι → α) [Fintype (PLift ι)] : Fintype (range f) := + Fintype.ofFinset (Finset.univ.image <| f ∘ PLift.down) <| by simp + +end FintypeInstances + +end Set + +/-! ### Finite instances + +There is seemingly some overlap between the following instances and the `Fintype` instances +in `Data.Set.Finite`. While every `Fintype` instance gives a `Finite` instance, those +instances that depend on `Fintype` or `Decidable` instances need an additional `Finite` instance +to be able to generally apply. + +Some set instances do not appear here since they are consequences of others, for example +`Subtype.Finite` for subsets of a finite type. +-/ + + +namespace Finite.Set + +open scoped Classical + +instance finite_range (f : ι → α) [Finite ι] : Finite (range f) := by + haveI := Fintype.ofFinite (PLift ι) + infer_instance + +instance finite_replacement [Finite α] (f : α → β) : + Finite {f x | x : α} := + Finite.Set.finite_range f + +end Finite.Set + +namespace Set + +/-! ### Constructors for `Set.Finite` + +Every constructor here should have a corresponding `Fintype` instance in the previous section +(or in the `Fintype` module). + +The implementation of these constructors ideally should be no more than `Set.toFinite`, +after possibly setting up some `Fintype` and classical `Decidable` instances. +-/ + + +section SetFiniteConstructors + +theorem finite_range (f : ι → α) [Finite ι] : (range f).Finite := + toFinite _ + +theorem Finite.dependent_image {s : Set α} (hs : s.Finite) (F : ∀ i ∈ s, β) : + {y : β | ∃ x hx, F x hx = y}.Finite := by + have := hs.to_subtype + simpa [range] using finite_range fun x : s => F x x.2 + +end SetFiniteConstructors + +end Set diff --git a/Mathlib/Data/Set/MemPartition.lean b/Mathlib/Data/Set/MemPartition.lean index 5a13d62ad7baf..94c1f4dca2984 100644 --- a/Mathlib/Data/Set/MemPartition.lean +++ b/Mathlib/Data/Set/MemPartition.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Data.Set.Finite +import Mathlib.Data.Set.Finite.Lattice /-! # Partitions based on membership of a sequence of sets diff --git a/Mathlib/Data/Set/Pointwise/Finite.lean b/Mathlib/Data/Set/Pointwise/Finite.lean index 23d93e1e23475..3b1d30164c535 100644 --- a/Mathlib/Data/Set/Pointwise/Finite.lean +++ b/Mathlib/Data/Set/Pointwise/Finite.lean @@ -5,7 +5,7 @@ Authors: Johan Commelin, Floris van Doorn -/ import Mathlib.Algebra.Group.Action.Basic import Mathlib.Algebra.Group.Pointwise.Set.Basic -import Mathlib.Data.Set.Finite +import Mathlib.Data.Finite.Prod /-! # Finiteness lemmas for pointwise operations on sets -/ diff --git a/Mathlib/Data/Setoid/Partition.lean b/Mathlib/Data/Setoid/Partition.lean index 27e668bcb2ece..c362e4ce974c2 100644 --- a/Mathlib/Data/Setoid/Partition.lean +++ b/Mathlib/Data/Setoid/Partition.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Amelia Livingston. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Amelia Livingston, Bryan Gin-ge Chen, Patrick Massot, Wen Yang, Johan Commelin -/ -import Mathlib.Data.Set.Finite +import Mathlib.Data.Set.Finite.Range import Mathlib.Order.Partition.Finpartition /-! diff --git a/Mathlib/GroupTheory/GroupAction/Basic.lean b/Mathlib/GroupTheory/GroupAction/Basic.lean index f394c0a339f82..ed5ab1ed24c4a 100644 --- a/Mathlib/GroupTheory/GroupAction/Basic.lean +++ b/Mathlib/GroupTheory/GroupAction/Basic.lean @@ -4,8 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ import Mathlib.Algebra.Group.Subgroup.Basic -import Mathlib.Data.Set.Finite +import Mathlib.Data.Finite.Sigma +import Mathlib.Data.Set.Finite.Basic +import Mathlib.Data.Set.Finite.Range import Mathlib.Data.Set.Pointwise.SMul +import Mathlib.Data.Setoid.Basic import Mathlib.GroupTheory.GroupAction.Defs /-! diff --git a/Mathlib/GroupTheory/GroupAction/Quotient.lean b/Mathlib/GroupTheory/GroupAction/Quotient.lean index 42ee1b3df994d..af70bfbcd8453 100644 --- a/Mathlib/GroupTheory/GroupAction/Quotient.lean +++ b/Mathlib/GroupTheory/GroupAction/Quotient.lean @@ -3,16 +3,15 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Thomas Browning -/ +import Mathlib.Algebra.Group.Subgroup.Actions +import Mathlib.Algebra.Group.Subgroup.ZPowers.Lemmas import Mathlib.Data.Fintype.BigOperators import Mathlib.Dynamics.PeriodicPts +import Mathlib.GroupTheory.Commutator.Basic +import Mathlib.GroupTheory.Coset.Basic import Mathlib.GroupTheory.GroupAction.Basic import Mathlib.GroupTheory.GroupAction.ConjAct import Mathlib.GroupTheory.GroupAction.Hom -import Mathlib.GroupTheory.Coset.Basic -import Mathlib.GroupTheory.Commutator.Basic -import Mathlib.Algebra.Group.Subgroup.Actions -import Mathlib.Data.Finite.Prod -import Mathlib.Algebra.Group.Subgroup.ZPowers.Lemmas /-! # Properties of group actions involving quotient groups diff --git a/Mathlib/GroupTheory/Perm/ClosureSwap.lean b/Mathlib/GroupTheory/Perm/ClosureSwap.lean index f30e18cc7099f..b00b89f6c6ce7 100644 --- a/Mathlib/GroupTheory/Perm/ClosureSwap.lean +++ b/Mathlib/GroupTheory/Perm/ClosureSwap.lean @@ -3,7 +3,6 @@ Copyright (c) 2024 Thomas Browning, Junyan Xu. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning, Junyan Xu -/ -import Mathlib.Data.Set.Finite import Mathlib.GroupTheory.GroupAction.Basic import Mathlib.GroupTheory.GroupAction.FixedPoints import Mathlib.GroupTheory.Perm.Support diff --git a/Mathlib/GroupTheory/Perm/Sign.lean b/Mathlib/GroupTheory/Perm/Sign.lean index 35402b7a4c405..d09dd86518f64 100644 --- a/Mathlib/GroupTheory/Perm/Sign.lean +++ b/Mathlib/GroupTheory/Perm/Sign.lean @@ -10,6 +10,7 @@ import Mathlib.Data.Finset.Fin import Mathlib.Data.Finset.Sort import Mathlib.Data.Fintype.Sum import Mathlib.Data.Int.Order.Units +import Mathlib.Data.Fintype.Prod import Mathlib.GroupTheory.Perm.Support import Mathlib.Logic.Equiv.Fin import Mathlib.Tactic.NormNum.Ineq diff --git a/Mathlib/GroupTheory/QuotientGroup/Finite.lean b/Mathlib/GroupTheory/QuotientGroup/Finite.lean index 6d234fdd50c16..d21f2e376e373 100644 --- a/Mathlib/GroupTheory/QuotientGroup/Finite.lean +++ b/Mathlib/GroupTheory/QuotientGroup/Finite.lean @@ -6,6 +6,7 @@ Authors: Kevin Buzzard, Patrick Massot -- This file is to a certain extent based on `quotient_module.lean` by Johannes Hölzl. import Mathlib.Algebra.Group.Subgroup.Finite +import Mathlib.Data.Finite.Prod import Mathlib.GroupTheory.QuotientGroup.Basic /-! diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index 9e75f4e16a58e..bef753458eb59 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -6,6 +6,7 @@ Authors: Sébastien Gouëzel import Mathlib.Algebra.Algebra.Defs import Mathlib.Algebra.NoZeroSMulDivisors.Pi import Mathlib.Data.Fintype.BigOperators +import Mathlib.Data.Fintype.Powerset import Mathlib.Data.Fintype.Sort import Mathlib.LinearAlgebra.Pi import Mathlib.Logic.Equiv.Fintype diff --git a/Mathlib/MeasureTheory/SetAlgebra.lean b/Mathlib/MeasureTheory/SetAlgebra.lean index d13040e7f87f2..168800fbfa2ed 100644 --- a/Mathlib/MeasureTheory/SetAlgebra.lean +++ b/Mathlib/MeasureTheory/SetAlgebra.lean @@ -3,6 +3,7 @@ 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.Data.Finite.Prod import Mathlib.MeasureTheory.SetSemiring /-! diff --git a/Mathlib/ModelTheory/FinitelyGenerated.lean b/Mathlib/ModelTheory/FinitelyGenerated.lean index 1ac2ffb332707..e8780fd26605b 100644 --- a/Mathlib/ModelTheory/FinitelyGenerated.lean +++ b/Mathlib/ModelTheory/FinitelyGenerated.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Aaron Anderson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ +import Mathlib.Data.Set.Finite.Lemmas import Mathlib.ModelTheory.Substructures /-! diff --git a/Mathlib/Order/Atoms/Finite.lean b/Mathlib/Order/Atoms/Finite.lean index 93cca39127a7d..dd0d8590f5a81 100644 --- a/Mathlib/Order/Atoms/Finite.lean +++ b/Mathlib/Order/Atoms/Finite.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Aaron Anderson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ +import Mathlib.Data.Set.Finite.Lattice import Mathlib.Order.Interval.Finset.Defs import Mathlib.Order.Atoms diff --git a/Mathlib/Order/Category/NonemptyFinLinOrd.lean b/Mathlib/Order/Category/NonemptyFinLinOrd.lean index 29ef2f1dc61d2..1bf3149bd7abe 100644 --- a/Mathlib/Order/Category/NonemptyFinLinOrd.lean +++ b/Mathlib/Order/Category/NonemptyFinLinOrd.lean @@ -3,14 +3,13 @@ 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.Data.Fintype.Order -import Mathlib.Data.Set.Finite -import Mathlib.Order.Category.FinPartOrd -import Mathlib.Order.Category.LinOrd +import Mathlib.CategoryTheory.ConcreteCategory.EpiMono import Mathlib.CategoryTheory.Limits.Shapes.Images import Mathlib.CategoryTheory.Limits.Shapes.RegularMono -import Mathlib.CategoryTheory.ConcreteCategory.EpiMono +import Mathlib.Data.Fintype.Order import Mathlib.Data.Set.Subsingleton +import Mathlib.Order.Category.FinPartOrd +import Mathlib.Order.Category.LinOrd /-! # Nonempty finite linear orders diff --git a/Mathlib/Order/CompleteSublattice.lean b/Mathlib/Order/CompleteSublattice.lean index 86da57c51a6d4..8eb523be561db 100644 --- a/Mathlib/Order/CompleteSublattice.lean +++ b/Mathlib/Order/CompleteSublattice.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Oliver Nash. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash -/ +import Mathlib.Data.Set.Functor import Mathlib.Order.Sublattice import Mathlib.Order.Hom.CompleteLattice diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Finset.lean b/Mathlib/Order/ConditionallyCompleteLattice/Finset.lean index 984c9aa53741f..3652f6c6faba4 100644 --- a/Mathlib/Order/ConditionallyCompleteLattice/Finset.lean +++ b/Mathlib/Order/ConditionallyCompleteLattice/Finset.lean @@ -3,8 +3,9 @@ Copyright (c) 2018 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ +import Mathlib.Data.Finset.Max +import Mathlib.Data.Set.Finite.Lattice import Mathlib.Order.ConditionallyCompleteLattice.Indexed -import Mathlib.Data.Set.Finite /-! # Conditionally complete lattices and finite sets. diff --git a/Mathlib/Order/Filter/AtTopBot.lean b/Mathlib/Order/Filter/AtTopBot.lean index c3666e61eba99..ed602722331c0 100644 --- a/Mathlib/Order/Filter/AtTopBot.lean +++ b/Mathlib/Order/Filter/AtTopBot.lean @@ -6,6 +6,7 @@ Authors: Johannes Hölzl, Jeremy Avigad, Yury Kudryashov, Patrick Massot import Mathlib.Data.Finset.Preimage import Mathlib.Order.ConditionallyCompleteLattice.Indexed import Mathlib.Order.Filter.Bases +import Mathlib.Data.Set.Finite.Lemmas import Mathlib.Order.Filter.Prod import Mathlib.Order.Interval.Set.Disjoint import Mathlib.Order.Interval.Set.OrderIso diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index ea155b49e6811..ee7fd636ed470 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Jeremy Avigad -/ -import Mathlib.Data.Set.Finite +import Mathlib.Data.Set.Finite.Lattice import Mathlib.Order.Filter.Defs /-! diff --git a/Mathlib/Order/Filter/Cofinite.lean b/Mathlib/Order/Filter/Cofinite.lean index a9b37d28a7653..50fbb43ffe460 100644 --- a/Mathlib/Order/Filter/Cofinite.lean +++ b/Mathlib/Order/Filter/Cofinite.lean @@ -3,6 +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, Jeremy Avigad, Yury Kudryashov -/ +import Mathlib.Data.Finite.Prod +import Mathlib.Data.Fintype.Pi import Mathlib.Order.Filter.AtTopBot import Mathlib.Order.Filter.Ker import Mathlib.Order.Filter.Pi diff --git a/Mathlib/Order/Interval/Set/Infinite.lean b/Mathlib/Order/Interval/Set/Infinite.lean index 9ab4354e14add..f6d1cfbc032d3 100644 --- a/Mathlib/Order/Interval/Set/Infinite.lean +++ b/Mathlib/Order/Interval/Set/Infinite.lean @@ -3,7 +3,8 @@ Copyright (c) 2020 Reid Barton. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Reid Barton -/ -import Mathlib.Data.Set.Finite +import Mathlib.Data.Set.Finite.Basic +import Mathlib.Order.Interval.Set.Basic /-! # Infinitude of intervals diff --git a/Mathlib/Order/KonigLemma.lean b/Mathlib/Order/KonigLemma.lean index 733e002d8ba39..a4ab85344bd14 100644 --- a/Mathlib/Order/KonigLemma.lean +++ b/Mathlib/Order/KonigLemma.lean @@ -5,6 +5,7 @@ Authors: Peter Nelson -/ import Mathlib.Order.Atoms.Finite import Mathlib.Order.Grade +import Mathlib.Tactic.ApplyFun /-! # Kőnig's infinity lemma diff --git a/Mathlib/Order/PartialSups.lean b/Mathlib/Order/PartialSups.lean index 41f07f8c01705..e4d11dbda28a8 100644 --- a/Mathlib/Order/PartialSups.lean +++ b/Mathlib/Order/PartialSups.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ import Mathlib.Data.Finset.Lattice.Fold -import Mathlib.Order.Hom.Basic -import Mathlib.Data.Set.Finite +import Mathlib.Data.Set.Finite.Lattice import Mathlib.Order.ConditionallyCompleteLattice.Indexed +import Mathlib.Order.Hom.Basic /-! # The monotone sequence of partial supremums of a sequence diff --git a/Mathlib/Order/Partition/Equipartition.lean b/Mathlib/Order/Partition/Equipartition.lean index 4cd1f1579f339..648a5421e9ba3 100644 --- a/Mathlib/Order/Partition/Equipartition.lean +++ b/Mathlib/Order/Partition/Equipartition.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.Order.Ring.Nat import Mathlib.Data.Set.Equitable import Mathlib.Logic.Equiv.Fin import Mathlib.Order.Partition.Finpartition +import Mathlib.Tactic.ApplyFun /-! # Finite equipartitions diff --git a/Mathlib/Order/Partition/Finpartition.lean b/Mathlib/Order/Partition/Finpartition.lean index 2753088bbea2a..e987fb0b84481 100644 --- a/Mathlib/Order/Partition/Finpartition.lean +++ b/Mathlib/Order/Partition/Finpartition.lean @@ -6,8 +6,8 @@ Authors: Yaël Dillies, Bhavik Mehta import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Data.Fintype.Powerset import Mathlib.Data.Setoid.Basic -import Mathlib.Order.SupIndep import Mathlib.Order.Atoms +import Mathlib.Order.SupIndep /-! # Finite partitions diff --git a/Mathlib/Order/SuccPred/LinearLocallyFinite.lean b/Mathlib/Order/SuccPred/LinearLocallyFinite.lean index e758ae8c0f2f4..c2dd78d6d45f1 100644 --- a/Mathlib/Order/SuccPred/LinearLocallyFinite.lean +++ b/Mathlib/Order/SuccPred/LinearLocallyFinite.lean @@ -8,6 +8,7 @@ import Mathlib.Logic.Encodable.Basic import Mathlib.Order.SuccPred.Archimedean import Mathlib.Order.Interval.Finset.Defs import Mathlib.Algebra.Order.Ring.Nat +import Mathlib.Data.Finset.Max /-! # Linear locally finite orders diff --git a/Mathlib/Order/SupClosed.lean b/Mathlib/Order/SupClosed.lean index 325063e488bda..61ddd8850360c 100644 --- a/Mathlib/Order/SupClosed.lean +++ b/Mathlib/Order/SupClosed.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Christopher Hoskin -/ import Mathlib.Data.Finset.Lattice.Fold -import Mathlib.Data.Set.Finite +import Mathlib.Data.Finset.Powerset +import Mathlib.Data.Set.Finite.Basic import Mathlib.Order.Closure /-! diff --git a/Mathlib/Order/UpperLower/LocallyFinite.lean b/Mathlib/Order/UpperLower/LocallyFinite.lean index 5c6d3e773f6cb..c0e733cf86794 100644 --- a/Mathlib/Order/UpperLower/LocallyFinite.lean +++ b/Mathlib/Order/UpperLower/LocallyFinite.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Yaël Dillies, Sara Rousta. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ +import Mathlib.Data.Set.Finite.Lattice import Mathlib.Order.Interval.Finset.Defs import Mathlib.Order.UpperLower.Basic diff --git a/Mathlib/RingTheory/Finiteness/Defs.lean b/Mathlib/RingTheory/Finiteness/Defs.lean index 31c1e84c16ca0..289071d1c7064 100644 --- a/Mathlib/RingTheory/Finiteness/Defs.lean +++ b/Mathlib/RingTheory/Finiteness/Defs.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ import Mathlib.Algebra.Algebra.Hom +import Mathlib.Data.Set.Finite.Lemmas import Mathlib.GroupTheory.Finiteness import Mathlib.RingTheory.Ideal.Span import Mathlib.Tactic.Algebraize diff --git a/Mathlib/RingTheory/Localization/Defs.lean b/Mathlib/RingTheory/Localization/Defs.lean index f92aa77e1af6f..486cd6a8eb03e 100644 --- a/Mathlib/RingTheory/Localization/Defs.lean +++ b/Mathlib/RingTheory/Localization/Defs.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Mario Carneiro, Johan Commelin, Amelia Livingston, Anne Baanen -/ +import Mathlib.Data.Fintype.Prod import Mathlib.GroupTheory.MonoidLocalization.MonoidWithZero import Mathlib.RingTheory.OreLocalization.Ring import Mathlib.Tactic.ApplyFun diff --git a/Mathlib/RingTheory/MvPolynomial/Basic.lean b/Mathlib/RingTheory/MvPolynomial/Basic.lean index 4f17802cf1e5f..3fc4ad5148f15 100644 --- a/Mathlib/RingTheory/MvPolynomial/Basic.lean +++ b/Mathlib/RingTheory/MvPolynomial/Basic.lean @@ -6,6 +6,7 @@ Authors: Johannes Hölzl import Mathlib.Algebra.CharP.Defs import Mathlib.Algebra.MvPolynomial.Degrees import Mathlib.Algebra.Polynomial.AlgebraMap +import Mathlib.Data.Fintype.Pi import Mathlib.LinearAlgebra.Finsupp.VectorSpace import Mathlib.LinearAlgebra.FreeModule.Finite.Basic diff --git a/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean b/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean index 2acf8eafd4ad5..e008a205dc742 100644 --- a/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean +++ b/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean @@ -11,6 +11,7 @@ import Mathlib.Algebra.GroupWithZero.Center import Mathlib.Algebra.Ring.Center import Mathlib.Algebra.Ring.Centralizer import Mathlib.Algebra.Ring.Prod +import Mathlib.Data.Set.Finite.Range import Mathlib.GroupTheory.Subsemigroup.Centralizer import Mathlib.RingTheory.NonUnitalSubsemiring.Defs diff --git a/Mathlib/RingTheory/Polynomial/Dickson.lean b/Mathlib/RingTheory/Polynomial/Dickson.lean index bf67e4ffe8272..d5d51a3dee789 100644 --- a/Mathlib/RingTheory/Polynomial/Dickson.lean +++ b/Mathlib/RingTheory/Polynomial/Dickson.lean @@ -238,7 +238,7 @@ theorem dickson_one_one_zmod_p (p : ℕ) [Fact p.Prime] : dickson 1 (1 : ZMod p) -- Finally, we prove the claim that our finite union of finite sets covers all of `K`. apply (Set.eq_univ_of_forall _).symm intro x - simp only [exists_prop, Set.mem_iUnion, Set.bind_def, Ne, Set.mem_setOf_eq] + simp only [exists_prop, Set.mem_iUnion, Ne, Set.mem_setOf_eq] by_cases hx : x = 0 · simp only [hx, and_true, eq_self_iff_true, inv_zero, or_true] exact ⟨_, 1, rfl, one_ne_zero⟩ diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index 0fd5a9930ac34..c3612fa0c43c8 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl, Mario Carneiro, Floris van Doorn -/ import Mathlib.Algebra.Order.Ring.Nat import Mathlib.Data.Fintype.BigOperators +import Mathlib.Data.Fintype.Powerset import Mathlib.Data.Nat.Cast.Order.Basic import Mathlib.Data.Set.Countable import Mathlib.Logic.Small.Set From 2c6b8bbf22415d2fad4a53b9f7f5b9047ad8077d Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Sat, 16 Nov 2024 08:30:46 +0000 Subject: [PATCH 012/829] feat: tensor product of two flat modules is flat (#18745) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add the missing `Flat R (M ⊗[R] N)` instance. Add two lemmas relating flatness to injectivity of `mapIncl`. Clean up the file Flat/Stability, renaming `Flat.comp` to `Flat.trans`. Also rename `Algebra.Flat.comp` and `FaithfullyFlat.comp`. --- .../LinearAlgebra/TensorProduct/Basic.lean | 4 ++ .../LinearAlgebra/TensorProduct/Tower.lean | 8 +++- Mathlib/RingTheory/Flat/Algebra.lean | 8 ++-- Mathlib/RingTheory/Flat/Basic.lean | 27 +++++++++++ Mathlib/RingTheory/Flat/FaithfullyFlat.lean | 13 +++--- Mathlib/RingTheory/Flat/Stability.lean | 45 ++++++++++--------- 6 files changed, 73 insertions(+), 32 deletions(-) diff --git a/Mathlib/LinearAlgebra/TensorProduct/Basic.lean b/Mathlib/LinearAlgebra/TensorProduct/Basic.lean index f849e4cda1696..231c9e14ba44f 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Basic.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Basic.lean @@ -1144,6 +1144,10 @@ lemma comm_comp_rTensor_comp_comm_eq (g : N →ₗ[R] P) : lTensor Q g := TensorProduct.ext rfl +theorem rTensor_tensor : rTensor (M ⊗[R] N) g = + TensorProduct.assoc R Q M N ∘ₗ rTensor N (rTensor M g) ∘ₗ (TensorProduct.assoc R P M N).symm := + TensorProduct.ext <| LinearMap.ext fun _ ↦ TensorProduct.ext rfl + lemma comm_comp_lTensor_comp_comm_eq (g : N →ₗ[R] P) : TensorProduct.comm R Q P ∘ₗ lTensor Q g ∘ₗ TensorProduct.comm R N Q = rTensor Q g := diff --git a/Mathlib/LinearAlgebra/TensorProduct/Tower.lean b/Mathlib/LinearAlgebra/TensorProduct/Tower.lean index cc69d012f40ef..f5564454ae805 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Tower.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Tower.lean @@ -338,8 +338,9 @@ variable [AddCommMonoid M] [Module R M] [Module A M] [Module B M] variable [IsScalarTower R A M] [IsScalarTower R B M] [SMulCommClass A B M] variable [AddCommMonoid N] [Module R N] variable [AddCommMonoid P] [Module A P] +variable [AddCommMonoid P'] [Module A P'] variable [AddCommMonoid Q] [Module R Q] -variable (R A B M N P Q) +variable (R A B M N P P' Q) attribute [local ext high] TensorProduct.ext @@ -372,6 +373,11 @@ theorem assoc_symm_tmul (m : M) (p : P) (q : Q) : (assoc R A B M P Q).symm (m ⊗ₜ (p ⊗ₜ q)) = (m ⊗ₜ p) ⊗ₜ q := rfl +theorem rTensor_tensor [Module R P'] [IsScalarTower R A P'] (g : P →ₗ[A] P') : + rTensor (M ⊗[R] N) g = + assoc R A A P' M N ∘ₗ map (rTensor M g) id ∘ₗ (assoc R A A P M N).symm.toLinearMap := + TensorProduct.ext <| LinearMap.ext fun _ ↦ ext fun _ _ ↦ rfl + end assoc section cancelBaseChange diff --git a/Mathlib/RingTheory/Flat/Algebra.lean b/Mathlib/RingTheory/Flat/Algebra.lean index 53bc49e2510e5..bf379a8806436 100644 --- a/Mathlib/RingTheory/Flat/Algebra.lean +++ b/Mathlib/RingTheory/Flat/Algebra.lean @@ -41,9 +41,11 @@ variable (R : Type u) (S : Type v) [CommRing R] [CommRing S] /-- If `T` is a flat `S`-algebra and `S` is a flat `R`-algebra, then `T` is a flat `R`-algebra. -/ -theorem comp (T : Type w) [CommRing T] [Algebra R S] [Algebra R T] [Algebra S T] +theorem trans (T : Type w) [CommRing T] [Algebra R S] [Algebra R T] [Algebra S T] [IsScalarTower R S T] [Algebra.Flat R S] [Algebra.Flat S T] : Algebra.Flat R T where - out := Module.Flat.comp R S T + out := Module.Flat.trans R S T + +@[deprecated (since := "2024-11-08")] alias comp := trans /-- If `S` is a flat `R`-algebra and `T` is any `R`-algebra, then `T ⊗[R] S` is a flat `T`-algebra. -/ @@ -78,6 +80,6 @@ variable {R : Type u} {S : Type v} {T : Type w} [CommRing R] [CommRing S] [CommR instance comp [RingHom.Flat f] [RingHom.Flat g] : RingHom.Flat (g.comp f) where out := by algebraize_only [f, g, g.comp f] - exact Algebra.Flat.comp R S T + exact Algebra.Flat.trans R S T end RingHom.Flat diff --git a/Mathlib/RingTheory/Flat/Basic.lean b/Mathlib/RingTheory/Flat/Basic.lean index caa8bd856a69e..9509f2f734fbb 100644 --- a/Mathlib/RingTheory/Flat/Basic.lean +++ b/Mathlib/RingTheory/Flat/Basic.lean @@ -260,6 +260,15 @@ theorem rTensor_preserves_injective_linearMap {N' : Type*} [AddCommGroup N'] [Mo @[deprecated (since := "2024-03-29")] alias preserves_injective_linearMap := rTensor_preserves_injective_linearMap +instance {S} [CommRing S] [Algebra R S] [Module S M] [IsScalarTower R S M] [Flat S M] [Flat R N] : + Flat S (M ⊗[R] N) := + (iff_rTensor_injective' _ _).mpr fun I ↦ by + simpa [AlgebraTensorModule.rTensor_tensor] using + rTensor_preserves_injective_linearMap (.restrictScalars R <| rTensor M I.subtype) + (rTensor_preserves_injective_linearMap _ I.injective_subtype) + +example [Flat R M] [Flat R N] : Flat R (M ⊗[R] N) := inferInstance + /-- If `M` is a flat module, then `𝟙 M ⊗ f` is injective for all injective linear maps `f`. -/ @@ -393,6 +402,24 @@ theorem iff_rTensor_exact : Flat R M ↔ Function.Exact f g → Function.Exact (f.rTensor M) (g.rTensor M) := iff_rTensor_exact'.{max u v} +variable (p : Submodule R M) (q : Submodule R N) + +/-- If p and q are submodules of M and N respectively, and M and q are flat, +then `p ⊗ q → M ⊗ N` is injective. -/ +theorem tensorProduct_mapIncl_injective_of_right + [Flat R M] [Flat R q] : Function.Injective (mapIncl p q) := by + rw [mapIncl, ← lTensor_comp_rTensor] + exact (lTensor_preserves_injective_linearMap _ q.injective_subtype).comp + (rTensor_preserves_injective_linearMap _ p.injective_subtype) + +/-- If p and q are submodules of M and N respectively, and N and p are flat, +then `p ⊗ q → M ⊗ N` is injective. -/ +theorem tensorProduct_mapIncl_injective_of_left + [Flat R p] [Flat R N] : Function.Injective (mapIncl p q) := by + rw [mapIncl, ← rTensor_comp_lTensor] + exact (rTensor_preserves_injective_linearMap _ p.injective_subtype).comp + (lTensor_preserves_injective_linearMap _ q.injective_subtype) + end Flat end Module diff --git a/Mathlib/RingTheory/Flat/FaithfullyFlat.lean b/Mathlib/RingTheory/Flat/FaithfullyFlat.lean index 737886d542ec6..203fd8a1cd50f 100644 --- a/Mathlib/RingTheory/Flat/FaithfullyFlat.lean +++ b/Mathlib/RingTheory/Flat/FaithfullyFlat.lean @@ -38,7 +38,7 @@ A module `M` over a commutative ring `R` is *faithfully flat* if it is flat and all linear maps `f : N → N'`, `f = 0` iff `f ⊗ M = 0`. - `Module.FaithfullyFlat.of_linearEquiv`: modules linearly equivalent to a flat modules are flat -- `Module.FaithfullyFlat.comp`: if `S` is `R`-faithfully flat and `M` is `S`-faithfully flat, then +- `Module.FaithfullyFlat.trans`: if `S` is `R`-faithfully flat and `M` is `S`-faithfully flat, then `M` is `R`-faithfully flat. - `Module.FaithfullyFlat.self`: the `R`-module `R` is faithfully flat. @@ -465,7 +465,7 @@ end fixed_universe end linearMap -section comp +section trans open TensorProduct LinearMap @@ -477,10 +477,9 @@ variable [FaithfullyFlat R S] [FaithfullyFlat S M] include S in /-- If `S` is a faithfully flat `R`-algebra, then any faithfully flat `S`-Module is faithfully flat as an `R`-module. -/ -theorem comp : - FaithfullyFlat R M := by +theorem trans : FaithfullyFlat R M := by rw [iff_zero_iff_lTensor_zero] - refine ⟨Module.Flat.comp R S M, @fun N _ _ N' _ _ f => ⟨fun aux => ?_, fun eq => eq ▸ by simp⟩⟩ + refine ⟨Module.Flat.trans R S M, @fun N _ _ N' _ _ f => ⟨fun aux => ?_, fun eq => eq ▸ by simp⟩⟩ rw [zero_iff_lTensor_zero (R:= R) (M := S) f, show f.lTensor S = (AlgebraTensorModule.map (A:= S) LinearMap.id f).restrictScalars R by aesop, show (0 : S ⊗[R] N →ₗ[R] S ⊗[R] N') = (0 : S ⊗[R] N →ₗ[S] S ⊗[R] N').restrictScalars R by rfl, @@ -489,7 +488,9 @@ theorem comp : apply_fun AlgebraTensorModule.cancelBaseChange R S S M N' using LinearEquiv.injective _ simpa using congr($aux (m ⊗ₜ[R] n)) -end comp +@[deprecated (since := "2024-11-08")] alias comp := trans + +end trans end FaithfullyFlat diff --git a/Mathlib/RingTheory/Flat/Stability.lean b/Mathlib/RingTheory/Flat/Stability.lean index 86f28d3004554..39fec41cfe501 100644 --- a/Mathlib/RingTheory/Flat/Stability.lean +++ b/Mathlib/RingTheory/Flat/Stability.lean @@ -62,13 +62,13 @@ private noncomputable abbrev auxRightMul (I : Ideal R) : M ⊗[R] I →ₗ[S] M private noncomputable abbrev J (I : Ideal R) : Ideal S := LinearMap.range (auxRightMul R S S I) -private noncomputable abbrev auxIso [Module.Flat R S] {I : Ideal R} : +private noncomputable abbrev auxIso [Flat R S] {I : Ideal R} : S ⊗[R] I ≃ₗ[S] J R S I := by apply LinearEquiv.ofInjective (auxRightMul R S S I) simp only [LinearMap.coe_comp, LinearEquiv.coe_coe, EquivLike.comp_injective] - exact (Module.Flat.iff_lTensor_injective' R S).mp inferInstance I + exact (Flat.iff_lTensor_injective' R S).mp inferInstance I -private noncomputable abbrev auxLTensor [Module.Flat R S] (I : Ideal R) : +private noncomputable abbrev auxLTensor [Flat R S] (I : Ideal R) : M ⊗[R] I →ₗ[S] M := by letI e1 : M ⊗[R] I ≃ₗ[S] M ⊗[S] (S ⊗[R] I) := (AlgebraTensorModule.cancelBaseChange R S S M I).symm @@ -78,24 +78,26 @@ private noncomputable abbrev auxLTensor [Module.Flat R S] (I : Ideal R) : letI e4 : M ⊗[S] S →ₗ[S] M := TensorProduct.rid S M exact e4 ∘ₗ e3 ∘ₗ (e1 ≪≫ₗ e2) -private lemma auxLTensor_eq [Module.Flat R S] {I : Ideal R} : +private lemma auxLTensor_eq [Flat R S] {I : Ideal R} : (auxLTensor R S M I : M ⊗[R] I →ₗ[R] M) = - TensorProduct.rid R M ∘ₗ lTensor M (I.subtype) := by + TensorProduct.rid R M ∘ₗ lTensor M I.subtype := by apply TensorProduct.ext' intro m x erw [TensorProduct.rid_tmul] simp /-- If `S` is a flat `R`-algebra, then any flat `S`-Module is also `R`-flat. -/ -theorem comp [Module.Flat R S] [Module.Flat S M] : Module.Flat R M := by - rw [Module.Flat.iff_lTensor_injective'] +theorem trans [Flat R S] [Flat S M] : Flat R M := by + rw [Flat.iff_lTensor_injective'] intro I rw [← EquivLike.comp_injective _ (TensorProduct.rid R M)] - haveI h : TensorProduct.rid R M ∘ lTensor M (Submodule.subtype I) = + haveI h : TensorProduct.rid R M ∘ lTensor M I.subtype = TensorProduct.rid R M ∘ₗ lTensor M I.subtype := rfl simp only [h, ← auxLTensor_eq R S M, LinearMap.coe_restrictScalars, LinearMap.coe_comp, LinearEquiv.coe_coe, EquivLike.comp_injective, EquivLike.injective_comp] - exact (Module.Flat.iff_lTensor_injective' S M).mp inferInstance _ + exact (Flat.iff_lTensor_injective' S M).mp inferInstance _ + +@[deprecated (since := "2024-11-03")] alias comp := trans end Composition @@ -124,27 +126,26 @@ private noncomputable abbrev auxRTensorBaseChange (I : Ideal S) : AlgebraTensorModule.cancelBaseChange R S S I M letI e2 : S ⊗[S] (S ⊗[R] M) ≃ₗ[S] S ⊗[R] M := AlgebraTensorModule.cancelBaseChange R S S S M - letI f : I ⊗[R] M →ₗ[S] S ⊗[R] M := AlgebraTensorModule.map - (Submodule.subtype I) LinearMap.id + letI f : I ⊗[R] M →ₗ[S] S ⊗[R] M := AlgebraTensorModule.map I.subtype LinearMap.id e2.symm.toLinearMap ∘ₗ f ∘ₗ e1.toLinearMap private lemma auxRTensorBaseChange_eq (I : Ideal S) : - auxRTensorBaseChange R S M I = rTensor (S ⊗[R] M) (Submodule.subtype I) := by + auxRTensorBaseChange R S M I = rTensor (S ⊗[R] M) I.subtype := by ext simp /-- If `M` is a flat `R`-module and `S` is any `R`-algebra, `S ⊗[R] M` is `S`-flat. -/ -instance baseChange [Module.Flat R M] : Module.Flat S (S ⊗[R] M) := by - rw [Module.Flat.iff_rTensor_injective'] +instance baseChange [Flat R M] : Flat S (S ⊗[R] M) := by + rw [Flat.iff_rTensor_injective'] intro I simp only [← auxRTensorBaseChange_eq, auxRTensorBaseChange, LinearMap.coe_comp, LinearEquiv.coe_coe, EmbeddingLike.comp_injective, EquivLike.injective_comp] exact rTensor_preserves_injective_linearMap (I.subtype : I →ₗ[R] S) Subtype.val_injective /-- A base change of a flat module is flat. -/ -theorem isBaseChange [Module.Flat R M] (N : Type t) [AddCommGroup N] [Module R N] [Module S N] +theorem isBaseChange [Flat R M] (N : Type t) [AddCommGroup N] [Module R N] [Module S N] [IsScalarTower R S N] {f : M →ₗ[R] N} (h : IsBaseChange S f) : - Module.Flat S N := + Flat S N := of_linearEquiv S (S ⊗[R] M) N (IsBaseChange.equiv h).symm end BaseChange @@ -155,16 +156,16 @@ variable {R : Type u} {M Mp : Type*} (Rp : Type v) [CommRing R] [AddCommGroup M] [Module R M] [CommRing Rp] [Algebra R Rp] [AddCommGroup Mp] [Module R Mp] [Module Rp Mp] [IsScalarTower R Rp Mp] -instance localizedModule [Module.Flat R M] (S : Submonoid R) : Module.Flat (Localization S) - (LocalizedModule S M) := by - apply Module.Flat.isBaseChange (R := R) (S := Localization S) +instance localizedModule [Flat R M] (S : Submonoid R) : + Flat (Localization S) (LocalizedModule S M) := by + apply Flat.isBaseChange (R := R) (S := Localization S) (f := LocalizedModule.mkLinearMap S M) rw [← isLocalizedModule_iff_isBaseChange S] exact localizedModuleIsLocalizedModule S -theorem of_isLocalizedModule [Module.Flat R M] (S : Submonoid R) [IsLocalization S Rp] - (f : M →ₗ[R] Mp) [h : IsLocalizedModule S f] : Module.Flat Rp Mp := by - fapply Module.Flat.isBaseChange (R := R) (M := M) (S := Rp) (N := Mp) +theorem of_isLocalizedModule [Flat R M] (S : Submonoid R) [IsLocalization S Rp] + (f : M →ₗ[R] Mp) [h : IsLocalizedModule S f] : Flat Rp Mp := by + fapply Flat.isBaseChange (R := R) (M := M) (S := Rp) (N := Mp) exact (isLocalizedModule_iff_isBaseChange S Rp f).mp h end Localization From bc7827342b0d7a3025f4a34100174550c59c4ca1 Mon Sep 17 00:00:00 2001 From: Alex Kontorovich <58564076+AlexKontorovich@users.noreply.github.com> Date: Sat, 16 Nov 2024 09:10:12 +0000 Subject: [PATCH 013/829] feat: `Complex.Rectangle` (#19052) This is another PR breaking #9598 into smaller pieces. We prove a variety of theorems about `reProdIm`, in particular defining a complex `Rectangle` and horizontal and vertical segments. --- Mathlib/Data/Complex/Basic.lean | 59 +++++++++++++++++++++++++++++++++ 1 file changed, 59 insertions(+) diff --git a/Mathlib/Data/Complex/Basic.lean b/Mathlib/Data/Complex/Basic.lean index 72098e77463e6..1762273d53338 100644 --- a/Mathlib/Data/Complex/Basic.lean +++ b/Mathlib/Data/Complex/Basic.lean @@ -839,4 +839,63 @@ unsafe instance instRepr : Repr ℂ where (if p > 65 then (Std.Format.bracket "(" · ")") else (·)) <| reprPrec f.re 65 ++ " + " ++ reprPrec f.im 70 ++ "*I" +section reProdIm + +/-- The preimage under `equivRealProd` of `s ×ˢ t` is `s ×ℂ t`. -/ +lemma preimage_equivRealProd_prod (s t : Set ℝ) : equivRealProd ⁻¹' (s ×ˢ t) = s ×ℂ t := rfl + +/-- The inequality `s × t ⊆ s₁ × t₁` holds in `ℂ` iff it holds in `ℝ × ℝ`. -/ +lemma reProdIm_subset_iff {s s₁ t t₁ : Set ℝ} : s ×ℂ t ⊆ s₁ ×ℂ t₁ ↔ s ×ˢ t ⊆ s₁ ×ˢ t₁ := by + rw [← @preimage_equivRealProd_prod s t, ← @preimage_equivRealProd_prod s₁ t₁] + exact Equiv.preimage_subset equivRealProd _ _ + +/-- If `s ⊆ s₁ ⊆ ℝ` and `t ⊆ t₁ ⊆ ℝ`, then `s × t ⊆ s₁ × t₁` in `ℂ`. -/ +lemma reProdIm_subset_iff' {s s₁ t t₁ : Set ℝ} : + s ×ℂ t ⊆ s₁ ×ℂ t₁ ↔ s ⊆ s₁ ∧ t ⊆ t₁ ∨ s = ∅ ∨ t = ∅ := by + convert prod_subset_prod_iff + exact reProdIm_subset_iff + +end reProdIm + +open scoped Interval + +section Rectangle + +/-- A `Rectangle` is an axis-parallel rectangle with corners `z` and `w`. -/ +def Rectangle (z w : ℂ) : Set ℂ := [[z.re, w.re]] ×ℂ [[z.im, w.im]] + +end Rectangle + +section Segments + +/-- A real segment `[a₁, a₂]` translated by `b * I` is the complex line segment. -/ +lemma horizontalSegment_eq (a₁ a₂ b : ℝ) : + (fun (x : ℝ) ↦ x + b * I) '' [[a₁, a₂]] = [[a₁, a₂]] ×ℂ {b} := by + rw [← preimage_equivRealProd_prod] + ext x + constructor + · intro hx + obtain ⟨x₁, hx₁, hx₁'⟩ := hx + simp [← hx₁', mem_preimage, mem_prod, hx₁] + · intro hx + obtain ⟨x₁, hx₁, hx₁', hx₁''⟩ := hx + refine ⟨x.re, x₁, by simp⟩ + +/-- A vertical segment `[b₁, b₂]` translated by `a` is the complex line segment. -/ +lemma verticalSegment_eq (a b₁ b₂ : ℝ) : + (fun (y : ℝ) ↦ a + y * I) '' [[b₁, b₂]] = {a} ×ℂ [[b₁, b₂]] := by + rw [← preimage_equivRealProd_prod] + ext x + constructor + · intro hx + obtain ⟨x₁, hx₁, hx₁'⟩ := hx + simp [← hx₁', mem_preimage, mem_prod, hx₁] + · intro hx + simp only [equivRealProd_apply, singleton_prod, mem_image, Prod.mk.injEq, + exists_eq_right_right, mem_preimage] at hx + obtain ⟨x₁, hx₁, hx₁', hx₁''⟩ := hx + refine ⟨x.im, x₁, by simp⟩ + +end Segments + end Complex From 6270f3467257beb5cf5118e21a98a0734356657e Mon Sep 17 00:00:00 2001 From: Yoh Tanimoto Date: Sat, 16 Nov 2024 09:33:49 +0000 Subject: [PATCH 014/829] feat(Topology/PartitionOfUnity): add variations of partition of unity for locally compact T2 space (#12266) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit add variations of `PartitionOfUnity` and Urysohn's lemma: - In a locally compact T2 space `X`, for a compact set `t` and a finite family of open sets `{s i}` such that `t ⊆ ⋃ i, s i`, there is a family of continuous function `{f i}` supported in `s i`, `∑ i, f i x = 1` on `t` and `0 ≤ f x ≤ 1`. - In a locally compact regular space `X`, for a compact set `t` and a closed set `s` that are disjoint, there is a continuous function `f` with compact support which is `1` on `t` and `0` on `s` and `f x Icc (0 : ℝ) 1 ` for all `x`. The former is formalized as a partition of unity. For this purpose, I extended `PartialRefiment` to include a predicate `p` about the refined set. With this, we can require that the refined open set has a compact closure. This can be applied when `X` is locally compact and T2. These variations are needed in order to prove that [rieszContentAux](https://leanprover-community.github.io/mathlib4_docs/Mathlib/MeasureTheory/Integral/RieszMarkovKakutani.html#rieszContentAux) is indeed a Borel measure (following Rudin "Real and complex analysis"), now in #18775. Co-authored-by: Yoh Tanimoto <57562556+yoh-tanimoto@users.noreply.github.com> --- Mathlib/Topology/PartitionOfUnity.lean | 100 +++++++++++++++++++++ Mathlib/Topology/ShrinkingLemma.lean | 119 +++++++++++++++++++++++++ Mathlib/Topology/UrysohnsLemma.lean | 32 ++++++- 3 files changed, 247 insertions(+), 4 deletions(-) diff --git a/Mathlib/Topology/PartitionOfUnity.lean b/Mathlib/Topology/PartitionOfUnity.lean index 1ff22bb263d6d..b78236ba907c2 100644 --- a/Mathlib/Topology/PartitionOfUnity.lean +++ b/Mathlib/Topology/PartitionOfUnity.lean @@ -315,6 +315,8 @@ instance : FunLike (BumpCovering ι X s) ι C(X, ℝ) where coe := toFun coe_injective' f g h := by cases f; cases g; congr +@[simp] lemma toFun_eq_coe : f.toFun = f := rfl + protected theorem locallyFinite : LocallyFinite fun i => support (f i) := f.locallyFinite' @@ -424,6 +426,51 @@ theorem exists_isSubordinate [NormalSpace X] [ParacompactSpace X] (hs : IsClosed rcases exists_isSubordinate_of_locallyFinite hs V hVo hVf hsV with ⟨f, hf⟩ exact ⟨f, hf.mono hVU⟩ +/-- If `X` is a locally compact T2 topological space and `U i`, `i : ι`, is a locally finite open +covering of a compact set `s`, then there exists a `BumpCovering ι X s` that is subordinate to `U`. +If `X` is a paracompact space, then the assumption `hf : LocallyFinite U` can be omitted, see +`BumpCovering.exists_isSubordinate`. This version assumes that `p : (X → ℝ) → Prop` is a predicate +that satisfies Urysohn's lemma, and provides a `BumpCovering` such that each function of the +covering satisfies `p`. -/ +theorem exists_isSubordinate_of_locallyFinite_of_prop_t2space [LocallyCompactSpace X] [T2Space X] + (p : (X → ℝ) → Prop) (h01 : ∀ s t, IsClosed s → IsCompact t → Disjoint s t → ∃ f : C(X, ℝ), + p f ∧ EqOn f 0 s ∧ EqOn f 1 t ∧ ∀ x, f x ∈ Icc (0 : ℝ) 1) + (hs : IsCompact s) (U : ι → Set X) (ho : ∀ i, IsOpen (U i)) (hf : LocallyFinite U) + (hU : s ⊆ ⋃ i, U i) : + ∃ f : BumpCovering ι X s, (∀ i, p (f i)) ∧ f.IsSubordinate U ∧ + ∀ i, HasCompactSupport (f i) := by + rcases exists_subset_iUnion_closure_subset_t2space hs ho (fun x _ => hf.point_finite x) hU with + ⟨V, hsV, hVo, hVU, hcp⟩ + have hVU' i : V i ⊆ U i := subset_closure.trans (hVU i) + rcases exists_subset_iUnion_closure_subset_t2space hs hVo + (fun x _ => (hf.subset hVU').point_finite x) hsV with ⟨W, hsW, hWo, hWV, hWc⟩ + choose f hfp hf0 hf1 hf01 using fun i => + h01 _ _ (isClosed_compl_iff.2 <| hVo i) (hWc i) + (disjoint_right.2 fun x hx => Classical.not_not.2 (hWV i hx)) + have hsupp i : support (f i) ⊆ V i := support_subset_iff'.2 (hf0 i) + refine ⟨⟨f, hf.subset fun i => Subset.trans (hsupp i) (hVU' i), fun i x => (hf01 i x).1, + fun i x => (hf01 i x).2, fun x hx => ?_⟩, + hfp, fun i => Subset.trans (closure_mono (hsupp i)) (hVU i), + fun i => IsCompact.of_isClosed_subset (hcp i) isClosed_closure <| closure_mono (hsupp i)⟩ + rcases mem_iUnion.1 (hsW hx) with ⟨i, hi⟩ + exact ⟨i, ((hf1 i).mono subset_closure).eventuallyEq_of_mem ((hWo i).mem_nhds hi)⟩ + +/-- If `X` is a normal topological space and `U i`, `i : ι`, is a locally finite open covering of a +closed set `s`, then there exists a `BumpCovering ι X s` that is subordinate to `U`. If `X` is a +paracompact space, then the assumption `hf : LocallyFinite U` can be omitted, see +`BumpCovering.exists_isSubordinate`. -/ +theorem exists_isSubordinate_hasCompactSupport_of_locallyFinite_t2space [LocallyCompactSpace X] + [T2Space X] + (hs : IsCompact s) (U : ι → Set X) (ho : ∀ i, IsOpen (U i)) (hf : LocallyFinite U) + (hU : s ⊆ ⋃ i, U i) : + ∃ f : BumpCovering ι X s, f.IsSubordinate U ∧ ∀ i, HasCompactSupport (f i) := by + -- need to switch 0 and 1 in `exists_continuous_zero_one_of_isCompact` + simpa using + exists_isSubordinate_of_locallyFinite_of_prop_t2space (fun _ => True) + (fun _ _ ht hs hd => + (exists_continuous_zero_one_of_isCompact' hs ht hd.symm).imp fun _ hf => ⟨trivial, hf⟩) + hs U ho hf hU + /-- Index of a bump function such that `fs i =ᶠ[𝓝 x] 1`. -/ def ind (x : X) (hx : x ∈ s) : ι := (f.eventuallyEq_one' x hx).choose @@ -576,4 +623,57 @@ theorem exists_isSubordinate [NormalSpace X] [ParacompactSpace X] (hs : IsClosed let ⟨f, hf⟩ := BumpCovering.exists_isSubordinate hs U ho hU ⟨f.toPartitionOfUnity, hf.toPartitionOfUnity⟩ +/-- If `X` is a locally compact T2 topological space and `U` is a locally finite open covering of a +compact set `s`, then there exists a `PartitionOfUnity ι X s` that is subordinate to `U`. -/ +theorem exists_isSubordinate_of_locallyFinite_t2space [LocallyCompactSpace X] [T2Space X] + (hs : IsCompact s) (U : ι → Set X) (ho : ∀ i, IsOpen (U i)) (hf : LocallyFinite U) + (hU : s ⊆ ⋃ i, U i) : + ∃ f : PartitionOfUnity ι X s, f.IsSubordinate U ∧ ∀ i, HasCompactSupport (f i) := + let ⟨f, hfsub, hfcp⟩ := + BumpCovering.exists_isSubordinate_hasCompactSupport_of_locallyFinite_t2space hs U ho hf hU + ⟨f.toPartitionOfUnity, hfsub.toPartitionOfUnity, fun i => IsCompact.of_isClosed_subset (hfcp i) + isClosed_closure <| closure_mono (f.support_toPartitionOfUnity_subset i)⟩ + end PartitionOfUnity + +/-- A variation of **Urysohn's lemma**. + +In a locally compact T2 space `X`, for a compact set `t` and a finite family of open sets `{s i}_i` +such that `t ⊆ ⋃ i, s i`, there is a family of compactly supported continuous functions `{f i}_i` +supported in `s i`, `∑ i, f i x = 1` on `t` and `0 ≤ f i x ≤ 1`. -/ +theorem exists_continuous_sum_one_of_isOpen_isCompact [T2Space X] [LocallyCompactSpace X] + {n : ℕ} {t : Set X} {s : Fin n → Set X} (hs : ∀ (i : Fin n), IsOpen (s i)) (htcp : IsCompact t) + (hst : t ⊆ ⋃ i, s i) : + ∃ f : Fin n → C(X, ℝ), (∀ (i : Fin n), tsupport (f i) ⊆ s i) ∧ EqOn (∑ i, f i) 1 t + ∧ (∀ (i : Fin n), ∀ (x : X), f i x ∈ Icc (0 : ℝ) 1) + ∧ (∀ (i : Fin n), HasCompactSupport (f i)) := by + obtain ⟨f, hfsub, hfcp⟩ := PartitionOfUnity.exists_isSubordinate_of_locallyFinite_t2space htcp s + hs (locallyFinite_of_finite _) hst + use f + refine ⟨fun i ↦ hfsub i, ?_, ?_, fun i => hfcp i⟩ + · intro x hx + simp only [Finset.sum_apply, Pi.one_apply] + have h := f.sum_eq_one' x hx + simp at h + rw [finsum_eq_sum (fun i => (f.toFun i) x) + (Finite.subset finite_univ (subset_univ (support fun i ↦ (f.toFun i) x)))] at h + simp only [Finite.toFinset_setOf, ne_eq] at h + rw [← h, ← Finset.sum_subset + (Finset.subset_univ (Finset.filter (fun (j : Fin n) ↦ ¬(f.toFun j) x = 0) Finset.univ)) + (by intro j hju hj + simp only [Finset.mem_filter, Finset.mem_univ, true_and, Decidable.not_not] at hj + exact hj)] + rfl + intro i x + refine ⟨f.nonneg i x, ?_⟩ + by_cases h0 : f i x = 0 + · rw [h0] + exact zero_le_one + rw [← Finset.sum_singleton (f · x) i] + apply le_trans _ (f.sum_le_one' x) + rw [finsum_eq_sum (f.toFun · x) (by exact toFinite (support (f.toFun · x)))] + simp only [Finite.toFinset_setOf, ne_eq] + gcongr with z hz + · exact fun j _ _ => f.nonneg j x + simp only [Finset.singleton_subset_iff, Finset.mem_filter, Finset.mem_univ, true_and] + exact h0 diff --git a/Mathlib/Topology/ShrinkingLemma.lean b/Mathlib/Topology/ShrinkingLemma.lean index eb77859f5152c..8f03eeab6a068 100644 --- a/Mathlib/Topology/ShrinkingLemma.lean +++ b/Mathlib/Topology/ShrinkingLemma.lean @@ -255,3 +255,122 @@ theorem exists_iUnion_eq_closed_subset (uo : ∀ i, IsOpen (u i)) (uf : ∀ x, { ⟨v, univ_subset_iff.1 vU, hv⟩ end NormalSpace + +section T2LocallyCompactSpace + +open ShrinkingLemma + +variable {u : ι → Set X} {s : Set X} [T2Space X] [LocallyCompactSpace X] + +/-- In a locally compact Hausdorff space `X`, if `s` is a compact set, `v` is a partial refinement, +and `i` is an index such that `i ∉ v.carrier`, then there exists a partial refinement that is +strictly greater than `v`. -/ +theorem exists_gt_t2space (v : PartialRefinement u s (fun w => IsCompact (closure w))) + (hs : IsCompact s) (i : ι) (hi : i ∉ v.carrier) : + ∃ v' : PartialRefinement u s (fun w => IsCompact (closure w)), + v < v' ∧ IsCompact (closure (v' i)) := by + -- take `v i` such that `closure (v i)` is compact + set si := s ∩ (⋃ j ≠ i, v j)ᶜ with hsi + simp only [ne_eq, compl_iUnion] at hsi + have hsic : IsCompact si := by + apply IsCompact.of_isClosed_subset hs _ Set.inter_subset_left + · have : IsOpen (⋃ j ≠ i, v j) := by + apply isOpen_biUnion + intro j _ + exact v.isOpen j + exact IsClosed.inter (IsCompact.isClosed hs) (IsOpen.isClosed_compl this) + have : si ⊆ v i := by + intro x hx + have (j) (hj : j ≠ i) : x ∉ v j := by + rw [hsi] at hx + apply Set.not_mem_of_mem_compl + have hsi' : x ∈ (⋂ i_1, ⋂ (_ : ¬i_1 = i), (v.toFun i_1)ᶜ) := Set.mem_of_mem_inter_right hx + rw [ne_eq] at hj + rw [Set.mem_iInter₂] at hsi' + exact hsi' j hj + obtain ⟨j, hj⟩ := Set.mem_iUnion.mp + (v.subset_iUnion (Set.mem_of_mem_inter_left hx)) + obtain rfl : j = i := by + by_contra! h + exact this j h hj + exact hj + obtain ⟨vi, hvi⟩ := exists_open_between_and_isCompact_closure hsic (v.isOpen i) this + classical + refine ⟨⟨update v i vi, insert i v.carrier, ?_, ?_, ?_, ?_, ?_⟩, ⟨?_, ?_⟩, ?_⟩ + · intro j + rcases eq_or_ne j i with (rfl| hne) <;> simp [*, v.isOpen] + · refine fun x hx => mem_iUnion.2 ?_ + rcases em (∃ j ≠ i, x ∈ v j) with (⟨j, hji, hj⟩ | h) + · use j + rwa [update_noteq hji] + · push_neg at h + use i + rw [update_same] + apply hvi.2.1 + rw [hsi] + exact ⟨hx, mem_iInter₂_of_mem h⟩ + · rintro j (rfl | hj) + · rw [update_same] + exact subset_trans hvi.2.2.1 <| PartialRefinement.subset v j + · rw [update_noteq (ne_of_mem_of_not_mem hj hi)] + exact v.closure_subset hj + · intro j hj + rw [mem_insert_iff] at hj + by_cases h : j = i + · rw [← h] + simp only [update_same] + exact hvi.2.2.2 + · apply hj.elim + · intro hji + exact False.elim (h hji) + · intro hjmemv + rw [update_noteq h] + exact v.pred_of_mem hjmemv + · intro j hj + rw [mem_insert_iff, not_or] at hj + rw [update_noteq hj.1, v.apply_eq hj.2] + · refine ⟨subset_insert _ _, fun j hj => ?_⟩ + exact (update_noteq (ne_of_mem_of_not_mem hj hi) _ _).symm + · exact fun hle => hi (hle.1 <| mem_insert _ _) + · simp only [update_same] + exact hvi.2.2.2 + +/-- **Shrinking lemma** . A point-finite open cover of a compact subset of a `T2Space` +`LocallyCompactSpace` can be "shrunk" to a new open cover so that the closure of each new open set +is contained in the corresponding original open set. -/ +theorem exists_subset_iUnion_closure_subset_t2space (hs : IsCompact s) (uo : ∀ i, IsOpen (u i)) + (uf : ∀ x ∈ s, { i | x ∈ u i }.Finite) (us : s ⊆ ⋃ i, u i) : + ∃ v : ι → Set X, s ⊆ iUnion v ∧ (∀ i, IsOpen (v i)) ∧ (∀ i, closure (v i) ⊆ u i) + ∧ (∀ i, IsCompact (closure (v i))) := by + haveI : Nonempty (PartialRefinement u s (fun w => IsCompact (closure w))) := + ⟨⟨u, ∅, uo, us, False.elim, False.elim, fun _ => rfl⟩⟩ + have : ∀ c : Set (PartialRefinement u s (fun w => IsCompact (closure w))), + IsChain (· ≤ ·) c → c.Nonempty → ∃ ub, ∀ v ∈ c, v ≤ ub := + fun c hc ne => ⟨.chainSup c hc ne uf us, fun v hv => PartialRefinement.le_chainSup _ _ _ _ hv⟩ + rcases zorn_le_nonempty this with ⟨v, hv⟩ + suffices ∀ i, i ∈ v.carrier from + ⟨v, v.subset_iUnion, fun i => v.isOpen _, fun i => v.closure_subset (this i), ?_⟩ + · intro i + exact v.pred_of_mem (this i) + · intro i + by_contra! hi + rcases exists_gt_t2space v hs i hi with ⟨v', hlt, _⟩ + exact hv.not_lt hlt + +/-- **Shrinking lemma**. A point-finite open cover of a compact subset of a locally compact T2 space +can be "shrunk" to a new closed cover so that each new closed set is contained in the corresponding +original open set. See also `exists_subset_iUnion_closure_subset_t2space` for a stronger statement. +-/ +theorem exists_subset_iUnion_compact_subset_t2space (hs : IsCompact s) (uo : ∀ i, IsOpen (u i)) + (uf : ∀ x ∈ s, { i | x ∈ u i }.Finite) (us : s ⊆ ⋃ i, u i) : + ∃ v : ι → Set X, s ⊆ iUnion v ∧ (∀ i, IsClosed (v i)) ∧ (∀ i, v i ⊆ u i) + ∧ ∀ i, IsCompact (v i) := by + let ⟨v, hsv, _, hv⟩ := exists_subset_iUnion_closure_subset_t2space hs uo uf us + use fun i => closure (v i) + refine ⟨?_, ?_, ?_⟩ + · exact Subset.trans hsv (iUnion_mono fun _ => subset_closure) + · simp only [isClosed_closure, implies_true] + · simp only + exact And.intro (fun i => hv.1 i) (fun i => hv.2 i) + +end T2LocallyCompactSpace diff --git a/Mathlib/Topology/UrysohnsLemma.lean b/Mathlib/Topology/UrysohnsLemma.lean index f645df48176ff..411f0e25ccac7 100644 --- a/Mathlib/Topology/UrysohnsLemma.lean +++ b/Mathlib/Topology/UrysohnsLemma.lean @@ -4,10 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Analysis.Normed.Affine.AddTorsor -import Mathlib.LinearAlgebra.AffineSpace.Ordered -import Mathlib.Topology.ContinuousMap.Basic import Mathlib.Analysis.NormedSpace.FunctionSeries import Mathlib.Analysis.SpecificLimits.Basic +import Mathlib.LinearAlgebra.AffineSpace.Ordered +import Mathlib.Topology.ContinuousMap.Algebra import Mathlib.Topology.GDelta.Basic /-! @@ -352,6 +352,30 @@ theorem exists_continuous_zero_one_of_isCompact [RegularSpace X] [LocallyCompact exact ⟨⟨c.lim, c.continuous_lim⟩, fun x hx ↦ c.lim_of_mem_C _ (sk.trans interior_subset hx), fun x hx => c.lim_of_nmem_U _ fun h => h hx, c.lim_mem_Icc⟩ +/-- Urysohn's lemma: if `s` and `t` are two disjoint sets in a regular locally compact topological +space `X`, with `s` compact and `t` closed, then there exists a continuous +function `f : X → ℝ` such that + +* `f` equals zero on `t`; +* `f` equals one on `s`; +* `0 ≤ f x ≤ 1` for all `x`. +-/ +theorem exists_continuous_zero_one_of_isCompact' [RegularSpace X] [LocallyCompactSpace X] + {s t : Set X} (hs : IsCompact s) (ht : IsClosed t) (hd : Disjoint s t) : + ∃ f : C(X, ℝ), EqOn f 0 t ∧ EqOn f 1 s ∧ ∀ x, f x ∈ Icc (0 : ℝ) 1 := by + obtain ⟨g, hgs, hgt, (hicc : ∀ x, 0 ≤ g x ∧ g x ≤ 1)⟩ := exists_continuous_zero_one_of_isCompact + hs ht hd + use 1 - g + refine ⟨?_, ?_, ?_⟩ + · intro x hx + simp only [ContinuousMap.sub_apply, ContinuousMap.one_apply, Pi.zero_apply] + exact sub_eq_zero_of_eq (id (EqOn.symm hgt) hx) + · intro x hx + simp only [ContinuousMap.sub_apply, ContinuousMap.one_apply, Pi.one_apply, sub_eq_self] + exact hgs hx + · intro x + simpa [and_comm] using hicc x + /-- Urysohn's lemma: if `s` and `t` are two disjoint sets in a regular locally compact topological space `X`, with `s` compact and `t` closed, then there exists a continuous compactly supported function `f : X → ℝ` such that @@ -438,8 +462,8 @@ theorem exists_continuous_one_zero_of_isCompact_of_isGδ [RegularSpace X] [Local compact open set `s` such that `t ⊆ s`, there is a continuous function `f` supported in `s`, `f x = 1` on `t` and `0 ≤ f x ≤ 1`. -/ lemma exists_tsupport_one_of_isOpen_isClosed [T2Space X] {s t : Set X} - (hs : IsOpen s) (hscp : IsCompact (closure s)) (ht : IsClosed t) (hst : t ⊆ s) : ∃ f : C(X, ℝ), - tsupport f ⊆ s ∧ EqOn f 1 t ∧ ∀ x, f x ∈ Icc (0 : ℝ) 1 := by + (hs : IsOpen s) (hscp : IsCompact (closure s)) (ht : IsClosed t) (hst : t ⊆ s) : + ∃ f : C(X, ℝ), tsupport f ⊆ s ∧ EqOn f 1 t ∧ ∀ x, f x ∈ Icc (0 : ℝ) 1 := by -- separate `sᶜ` and `t` by `u` and `v`. rw [← compl_compl s] at hscp obtain ⟨u, v, huIsOpen, hvIsOpen, hscompl_subset_u, ht_subset_v, hDjsjointuv⟩ := From f974247f3d98dc3556a307d04f2db50fd12b753b Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sat, 16 Nov 2024 15:30:36 +0000 Subject: [PATCH 015/829] feat(KrullDimension): height_eq_coe_iff_minimal_le_height (#19110) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The elements of finite height `n` are the minimial elements among those of height `≥ n`. From the Carleson project --- Mathlib/Order/KrullDimension.lean | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/Mathlib/Order/KrullDimension.lean b/Mathlib/Order/KrullDimension.lean index 3a6129a85415f..da2a6a15b28c5 100644 --- a/Mathlib/Order/KrullDimension.lean +++ b/Mathlib/Order/KrullDimension.lean @@ -5,6 +5,7 @@ Authors: Jujian Zhang, Fangming Li, Joachim Breitner -/ import Mathlib.Order.RelSeries +import Mathlib.Order.Minimal import Mathlib.Data.ENat.Lattice /-! @@ -318,6 +319,21 @@ lemma height_eq_coe_iff {x : α} {n : ℕ} : rename_i y _ cases height y <;> simp; norm_cast; omega +/-- The elements of finite height `n` are the minimial elements among those of height `≥ n`. -/ +lemma height_eq_coe_iff_minimal_le_height {a : α} {n : ℕ} : + height a = n ↔ Minimal (fun y => n ≤ height y) a := by + by_cases hfin : height a < ⊤ + · cases hn : n with + | zero => simp + | succ => simp [minimal_iff_forall_lt, height_eq_coe_add_one_iff, ENat.add_one_le_iff, + coe_lt_height_iff, *] + · suffices ∃ x < a, ↑n ≤ height x by + simp_all [minimal_iff_forall_lt] + simp only [not_lt, top_le_iff, height_eq_top_iff] at hfin + obtain ⟨p, rfl, hp⟩ := hfin (n+1) + use p.eraseLast.last, RelSeries.eraseLast_last_rel_last _ (by omega) + simpa [hp] using length_le_height_last (p := p.eraseLast) + end height /-! From b4ef11754f83859f04237acbd0c2375b8cb0e354 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Sat, 16 Nov 2024 15:42:22 +0000 Subject: [PATCH 016/829] chore: don't use `TopologicalGroup.toUniformSpace` in `tendstoUniformly_iff` (#19098) This changes `TopologicalGroup.tendstoUniformly_iff` so that it takes a `UniformSpace` instance instead of only a `TopologicalGroup` instance, along with a proof that the uniform space is the one induced by the topological group structure. The previous version would have been hard to use in case you had, for example, a normed commutative group, as the uniform space structure inferred would not match (defeq) the one given in this theorem, which was previously `TopologicalSpace.toUniformSpace`. --- .../Topology/Algebra/UniformGroup/Basic.lean | 48 ++++++++++--------- 1 file changed, 26 insertions(+), 22 deletions(-) diff --git a/Mathlib/Topology/Algebra/UniformGroup/Basic.lean b/Mathlib/Topology/Algebra/UniformGroup/Basic.lean index 6d99d5ffcee64..8fd7dd61a29d2 100644 --- a/Mathlib/Topology/Algebra/UniformGroup/Basic.lean +++ b/Mathlib/Topology/Algebra/UniformGroup/Basic.lean @@ -207,37 +207,41 @@ lemma MonoidHom.tendsto_coe_cofinite_of_discrete [T2Space G] {H : Type*} [Group replace hf : Function.Injective f.rangeRestrict := by simpa exact f.range.tendsto_coe_cofinite_of_discrete.comp hf.tendsto_cofinite +end TopologicalGroup + +namespace TopologicalGroup + +variable {ι α G : Type*} [Group G] [u : UniformSpace G] [TopologicalGroup G] + @[to_additive] -theorem TopologicalGroup.tendstoUniformly_iff {ι α : Type*} (F : ι → α → G) (f : α → G) - (p : Filter ι) : - @TendstoUniformly α G ι (TopologicalGroup.toUniformSpace G) F f p ↔ - ∀ u ∈ 𝓝 (1 : G), ∀ᶠ i in p, ∀ a, F i a / f a ∈ u := - ⟨fun h u hu => h _ ⟨u, hu, fun _ => id⟩, fun h _ ⟨u, hu, hv⟩ => - mem_of_superset (h u hu) fun _ hi a => hv (hi a)⟩ +theorem tendstoUniformly_iff (F : ι → α → G) (f : α → G) (p : Filter ι) + (hu : TopologicalGroup.toUniformSpace G = u) : + TendstoUniformly F f p ↔ ∀ u ∈ 𝓝 (1 : G), ∀ᶠ i in p, ∀ a, F i a / f a ∈ u := + hu ▸ ⟨fun h u hu => h _ ⟨u, hu, fun _ => id⟩, + fun h _ ⟨u, hu, hv⟩ => mem_of_superset (h u hu) fun _ hi a => hv (hi a)⟩ @[to_additive] -theorem TopologicalGroup.tendstoUniformlyOn_iff {ι α : Type*} (F : ι → α → G) (f : α → G) - (p : Filter ι) (s : Set α) : - @TendstoUniformlyOn α G ι (TopologicalGroup.toUniformSpace G) F f p s ↔ - ∀ u ∈ 𝓝 (1 : G), ∀ᶠ i in p, ∀ a ∈ s, F i a / f a ∈ u := - ⟨fun h u hu => h _ ⟨u, hu, fun _ => id⟩, fun h _ ⟨u, hu, hv⟩ => - mem_of_superset (h u hu) fun _ hi a ha => hv (hi a ha)⟩ +theorem tendstoUniformlyOn_iff (F : ι → α → G) (f : α → G) (p : Filter ι) (s : Set α) + (hu : TopologicalGroup.toUniformSpace G = u) : + TendstoUniformlyOn F f p s ↔ ∀ u ∈ 𝓝 (1 : G), ∀ᶠ i in p, ∀ a ∈ s, F i a / f a ∈ u := + hu ▸ ⟨fun h u hu => h _ ⟨u, hu, fun _ => id⟩, + fun h _ ⟨u, hu, hv⟩ => mem_of_superset (h u hu) fun _ hi a ha => hv (hi a ha)⟩ @[to_additive] -theorem TopologicalGroup.tendstoLocallyUniformly_iff {ι α : Type*} [TopologicalSpace α] - (F : ι → α → G) (f : α → G) (p : Filter ι) : - @TendstoLocallyUniformly α G ι (TopologicalGroup.toUniformSpace G) _ F f p ↔ +theorem tendstoLocallyUniformly_iff [TopologicalSpace α] (F : ι → α → G) (f : α → G) + (p : Filter ι) (hu : TopologicalGroup.toUniformSpace G = u) : + TendstoLocallyUniformly F f p ↔ ∀ u ∈ 𝓝 (1 : G), ∀ (x : α), ∃ t ∈ 𝓝 x, ∀ᶠ i in p, ∀ a ∈ t, F i a / f a ∈ u := - ⟨fun h u hu => h _ ⟨u, hu, fun _ => id⟩, fun h _ ⟨u, hu, hv⟩ x => - Exists.imp (fun _ ⟨h, hp⟩ => ⟨h, mem_of_superset hp fun _ hi a ha => hv (hi a ha)⟩) - (h u hu x)⟩ + hu ▸ ⟨fun h u hu => h _ ⟨u, hu, fun _ => id⟩, fun h _ ⟨u, hu, hv⟩ x => + Exists.imp (fun _ ⟨h, hp⟩ => ⟨h, mem_of_superset hp fun _ hi a ha => hv (hi a ha)⟩) + (h u hu x)⟩ @[to_additive] -theorem TopologicalGroup.tendstoLocallyUniformlyOn_iff {ι α : Type*} [TopologicalSpace α] - (F : ι → α → G) (f : α → G) (p : Filter ι) (s : Set α) : - @TendstoLocallyUniformlyOn α G ι (TopologicalGroup.toUniformSpace G) _ F f p s ↔ +theorem tendstoLocallyUniformlyOn_iff [TopologicalSpace α] (F : ι → α → G) (f : α → G) + (p : Filter ι) (s : Set α) (hu : TopologicalGroup.toUniformSpace G = u) : + TendstoLocallyUniformlyOn F f p s ↔ ∀ u ∈ 𝓝 (1 : G), ∀ x ∈ s, ∃ t ∈ 𝓝[s] x, ∀ᶠ i in p, ∀ a ∈ t, F i a / f a ∈ u := - ⟨fun h u hu => h _ ⟨u, hu, fun _ => id⟩, fun h _ ⟨u, hu, hv⟩ x => + hu ▸ ⟨fun h u hu => h _ ⟨u, hu, fun _ => id⟩, fun h _ ⟨u, hu, hv⟩ x => (Exists.imp fun _ ⟨h, hp⟩ => ⟨h, mem_of_superset hp fun _ hi a ha => hv (hi a ha)⟩) ∘ h u hu x⟩ From 7b8edbb98887af54a4cb8e8f76fa279960233c22 Mon Sep 17 00:00:00 2001 From: JonBannon Date: Sat, 16 Nov 2024 16:13:10 +0000 Subject: [PATCH 017/829] chore(Algebra.Order.Star.Basic): Protected `IsSelfAdjoint.mul_self_nonneg` and `IsSelfAdjoint.sq_nonneg` (#19128) Marking these lemmas as `protected`. --- Mathlib/Algebra/Order/Star/Basic.lean | 4 ++-- .../CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean | 4 ++-- 2 files changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/Algebra/Order/Star/Basic.lean b/Mathlib/Algebra/Order/Star/Basic.lean index 0297489633917..71d84e9eeb151 100644 --- a/Mathlib/Algebra/Order/Star/Basic.lean +++ b/Mathlib/Algebra/Order/Star/Basic.lean @@ -169,7 +169,7 @@ theorem mul_star_self_nonneg (r : R) : 0 ≤ r * star r := by simpa only [star_star] using star_mul_self_nonneg (star r) @[aesop safe apply (rule_sets := [CStarAlgebra])] -theorem IsSelfAdjoint.mul_self_nonneg {a : R} (ha : IsSelfAdjoint a) : 0 ≤ a * a := by +protected theorem IsSelfAdjoint.mul_self_nonneg {a : R} (ha : IsSelfAdjoint a) : 0 ≤ a * a := by simpa [ha.star_eq] using star_mul_self_nonneg a @[aesop safe apply] @@ -302,7 +302,7 @@ lemma star_lt_one_iff {x : R} : star x < 1 ↔ x < 1 := by simpa using star_lt_star_iff (x := x) (y := 1) @[aesop safe apply (rule_sets := [CStarAlgebra])] -theorem IsSelfAdjoint.sq_nonneg {a : R} (ha : IsSelfAdjoint a) : 0 ≤ a ^ 2 := by +protected theorem IsSelfAdjoint.sq_nonneg {a : R} (ha : IsSelfAdjoint a) : 0 ≤ a ^ 2 := by simp [sq, ha.mul_self_nonneg] end Semiring diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean index b50bedc7fc13c..76c6e66606ed2 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean @@ -409,11 +409,11 @@ lemma SpectrumRestricts.nnreal_add {a b : A} (ha₁ : IsSelfAdjoint a) gcongr all_goals rw [← SpectrumRestricts.nnreal_iff_nnnorm] <;> first | rfl | assumption -protected lemma IsSelfAdjoint.sq_spectrumRestricts {a : A} (ha : IsSelfAdjoint a) : +lemma IsSelfAdjoint.sq_spectrumRestricts {a : A} (ha : IsSelfAdjoint a) : SpectrumRestricts (a ^ 2) ContinuousMap.realToNNReal := by rw [SpectrumRestricts.nnreal_iff, ← cfc_id (R := ℝ) a, ← cfc_pow .., cfc_map_spectrum ..] rintro - ⟨x, -, rfl⟩ - exact _root_.sq_nonneg (id x) + exact sq_nonneg x open ComplexStarModule From 2aeb15532ae2dcb0c4119ce6e5e5825811af14f1 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Sat, 16 Nov 2024 20:53:18 +0000 Subject: [PATCH 018/829] feat(NumberTheory/ModularForms): q-expansions of modular forms (#18813) Show that modular forms of level `n` can be written as analytic functions of `q = exp (2 * pi * I * z / n)`. Co-authored-by: Chris Birkbeck --- Mathlib.lean | 1 + Mathlib/Analysis/Complex/Periodic.lean | 6 +- .../Complex/UpperHalfPlane/Basic.lean | 5 + .../FunctionsBoundedAtInfty.lean | 15 ++- .../Complex/UpperHalfPlane/Manifold.lean | 40 +++++- .../Complex/UpperHalfPlane/Topology.lean | 16 ++- .../NumberTheory/ModularForms/QExpansion.lean | 125 ++++++++++++++++++ 7 files changed, 202 insertions(+), 6 deletions(-) create mode 100644 Mathlib/NumberTheory/ModularForms/QExpansion.lean diff --git a/Mathlib.lean b/Mathlib.lean index 679585a362a5a..ef683fa1d3328 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3782,6 +3782,7 @@ import Mathlib.NumberTheory.ModularForms.JacobiTheta.Bounds import Mathlib.NumberTheory.ModularForms.JacobiTheta.Manifold import Mathlib.NumberTheory.ModularForms.JacobiTheta.OneVariable import Mathlib.NumberTheory.ModularForms.JacobiTheta.TwoVariable +import Mathlib.NumberTheory.ModularForms.QExpansion import Mathlib.NumberTheory.ModularForms.SlashActions import Mathlib.NumberTheory.ModularForms.SlashInvariantForms import Mathlib.NumberTheory.MulChar.Basic diff --git a/Mathlib/Analysis/Complex/Periodic.lean b/Mathlib/Analysis/Complex/Periodic.lean index 6267fdbc4d0b8..2237eb83d967e 100644 --- a/Mathlib/Analysis/Complex/Periodic.lean +++ b/Mathlib/Analysis/Complex/Periodic.lean @@ -20,7 +20,7 @@ for all sufficiently large `im z`, then `F` extends to a holomorphic function on to zero. These results are important in the theory of modular forms. -/ -open Complex Filter Asymptotics Function +open Complex Filter Asymptotics open scoped Real Topology @@ -30,6 +30,8 @@ local notation "I∞" => comap im atTop variable (h : ℝ) +namespace Function.Periodic + /-- Parameter for q-expansions, `qParam h z = exp (2 * π * I * z / h)` -/ def qParam (z : ℂ) : ℂ := exp (2 * π * I * z / h) @@ -214,3 +216,5 @@ theorem exp_decay_of_zero_at_inf (hh : 0 < h) (hf : Periodic f h) nhdsWithin_le_nhds end HoloAtInfC + +end Function.Periodic diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean index 1ae8bc92b4330..52bf76205cdd1 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean @@ -104,6 +104,11 @@ theorem mk_im (z : ℂ) (h : 0 < z.im) : (mk z h).im = z.im := theorem coe_mk (z : ℂ) (h : 0 < z.im) : (mk z h : ℂ) = z := rfl +@[simp] +lemma coe_mk_subtype {z : ℂ} (hz : 0 < z.im) : + UpperHalfPlane.coe ⟨z, hz⟩ = z := by + rfl + @[simp] theorem mk_coe (z : ℍ) (h : 0 < (z : ℂ).im := z.2) : mk z h = z := rfl diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/FunctionsBoundedAtInfty.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/FunctionsBoundedAtInfty.lean index 433fc78c1f9e5..7ecb4b7001f40 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/FunctionsBoundedAtInfty.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/FunctionsBoundedAtInfty.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Birkbeck, David Loeffler -/ import Mathlib.Algebra.Module.Submodule.Basic -import Mathlib.Analysis.Complex.UpperHalfPlane.Basic +import Mathlib.Analysis.Complex.UpperHalfPlane.Topology import Mathlib.Order.Filter.ZeroAndBoundedAtFilter /-! @@ -67,4 +67,17 @@ theorem zero_at_im_infty (f : ℍ → ℂ) : (atImInfty_basis.tendsto_iff Metric.nhds_basis_closedBall).trans <| by simp only [true_and, mem_closedBall_zero_iff]; rfl +lemma tendsto_comap_im_ofComplex : + Tendsto ofComplex (comap Complex.im atTop) atImInfty := by + simp only [atImInfty, tendsto_comap_iff, Function.comp_def] + refine tendsto_comap.congr' ?_ + filter_upwards [preimage_mem_comap (Ioi_mem_atTop 0)] with z hz + simp only [ofComplex_apply_of_im_pos hz, ← UpperHalfPlane.coe_im, coe_mk_subtype] + +lemma tendsto_coe_atImInfty : + Tendsto UpperHalfPlane.coe atImInfty (comap Complex.im atTop) := by + simpa only [atImInfty, tendsto_comap_iff, Function.comp_def, + funext UpperHalfPlane.coe_im] using tendsto_comap + + end UpperHalfPlane diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Manifold.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Manifold.lean index aede75f666ab9..c821c8feafdcb 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Manifold.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Manifold.lean @@ -1,11 +1,11 @@ /- Copyright (c) 2022 Chris Birkbeck. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Chris Birkbeck +Authors: Chris Birkbeck, David Loeffler -/ import Mathlib.Analysis.Complex.UpperHalfPlane.Topology import Mathlib.Geometry.Manifold.ContMDiff.Atlas -import Mathlib.Geometry.Manifold.MFDeriv.Basic +import Mathlib.Geometry.Manifold.MFDeriv.FDeriv /-! # Manifold structure on the upper half plane. @@ -13,8 +13,9 @@ import Mathlib.Geometry.Manifold.MFDeriv.Basic In this file we define the complex manifold structure on the upper half-plane. -/ +open Filter -open scoped UpperHalfPlane Manifold +open scoped Manifold namespace UpperHalfPlane @@ -31,4 +32,37 @@ theorem smooth_coe : Smooth 𝓘(ℂ) 𝓘(ℂ) ((↑) : ℍ → ℂ) := fun _ = theorem mdifferentiable_coe : MDifferentiable 𝓘(ℂ) 𝓘(ℂ) ((↑) : ℍ → ℂ) := smooth_coe.mdifferentiable +lemma smoothAt_ofComplex {z : ℂ} (hz : 0 < z.im) : + SmoothAt 𝓘(ℂ) 𝓘(ℂ) ofComplex z := by + rw [SmoothAt, contMDiffAt_iff] + constructor + · -- continuity at z + rw [ContinuousAt, nhds_induced, tendsto_comap_iff] + refine Tendsto.congr' (eventuallyEq_coe_comp_ofComplex hz).symm ?_ + simpa only [ofComplex_apply_of_im_pos hz, Subtype.coe_mk] using tendsto_id + · -- smoothness in local chart + simp only [extChartAt, PartialHomeomorph.extend, modelWithCornersSelf_partialEquiv, + PartialEquiv.trans_refl, PartialHomeomorph.toFun_eq_coe, PartialHomeomorph.refl_partialEquiv, + PartialEquiv.refl_source, PartialHomeomorph.singletonChartedSpace_chartAt_eq, + PartialEquiv.refl_symm, PartialEquiv.refl_coe, CompTriple.comp_eq, modelWithCornersSelf_coe, + Set.range_id, id_eq, contDiffWithinAt_univ] + exact contDiffAt_id.congr_of_eventuallyEq (eventuallyEq_coe_comp_ofComplex hz) + +lemma mdifferentiableAt_ofComplex {z : ℂ} (hz : 0 < z.im) : + MDifferentiableAt 𝓘(ℂ) 𝓘(ℂ) ofComplex z := + (smoothAt_ofComplex hz).mdifferentiableAt + +lemma mdifferentiableAt_iff {f : ℍ → ℂ} {τ : ℍ} : + MDifferentiableAt 𝓘(ℂ) 𝓘(ℂ) f τ ↔ DifferentiableAt ℂ (f ∘ ofComplex) ↑τ := by + rw [← mdifferentiableAt_iff_differentiableAt] + refine ⟨fun hf ↦ ?_, fun hf ↦ ?_⟩ + · exact (ofComplex_apply τ ▸ hf).comp _ (mdifferentiableAt_ofComplex τ.im_pos) + · simpa only [Function.comp_def, ofComplex_apply] using hf.comp τ (mdifferentiable_coe τ) + +lemma mdifferentiable_iff {f : ℍ → ℂ} : + MDifferentiable 𝓘(ℂ) 𝓘(ℂ) f ↔ DifferentiableOn ℂ (f ∘ ofComplex) {z | 0 < z.im} := + ⟨fun h z hz ↦ (mdifferentiableAt_iff.mp (h ⟨z, hz⟩)).differentiableWithinAt, + fun h ⟨z, hz⟩ ↦ mdifferentiableAt_iff.mpr <| (h z hz).differentiableAt + <| (Complex.continuous_im.isOpen_preimage _ isOpen_Ioi).mem_nhds hz⟩ + end UpperHalfPlane diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean index 2187ae0e7a527..ac416004b29f9 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean @@ -121,12 +121,15 @@ theorem ModularGroup_T_zpow_mem_verticalStrip (z : ℍ) {N : ℕ} (hn : 0 < N) : end strips -/-- A continuous section `ℂ → ℍ` of the natural inclusion map, bundled as a `PartialHomeomorph`. -/ +section ofComplex + +/-- A section `ℂ → ℍ` of the natural inclusion map, bundled as a `PartialHomeomorph`. -/ def ofComplex : PartialHomeomorph ℂ ℍ := (isOpenEmbedding_coe.toPartialHomeomorph _).symm /-- Extend a function on `ℍ` arbitrarily to a function on all of `ℂ`. -/ scoped notation "↑ₕ" f => f ∘ ofComplex +@[simp] lemma ofComplex_apply (z : ℍ) : ofComplex (z : ℂ) = z := IsOpenEmbedding.toPartialHomeomorph_left_inv .. @@ -139,6 +142,10 @@ lemma ofComplex_apply_eq_ite (w : ℂ) : rintro ⟨a, rfl⟩ exact (a.prop.not_le (by simpa using hw)).elim +lemma ofComplex_apply_of_im_pos {z : ℂ} (hz : 0 < z.im) : + ofComplex z = ⟨z, hz⟩ := by + simpa only [coe_mk_subtype] using ofComplex_apply ⟨z, hz⟩ + lemma ofComplex_apply_of_im_nonpos {w : ℂ} (hw : w.im ≤ 0) : ofComplex w = Classical.choice inferInstance := by simp [ofComplex_apply_eq_ite w, hw] @@ -157,4 +164,11 @@ lemma comp_ofComplex_of_im_le_zero (f : ℍ → ℂ) (z z' : ℂ) (hz : z.im ≤ (↑ₕ f) z = (↑ₕ f) z' := by simp [ofComplex_apply_of_im_nonpos, hz, hz'] +lemma eventuallyEq_coe_comp_ofComplex {z : ℂ} (hz : 0 < z.im) : + UpperHalfPlane.coe ∘ ofComplex =ᶠ[𝓝 z] id := by + filter_upwards [(Complex.continuous_im.isOpen_preimage _ isOpen_Ioi).mem_nhds hz] with x hx + simp only [Function.comp_apply, ofComplex_apply_of_im_pos hx, id_eq, coe_mk_subtype] + +end ofComplex + end UpperHalfPlane diff --git a/Mathlib/NumberTheory/ModularForms/QExpansion.lean b/Mathlib/NumberTheory/ModularForms/QExpansion.lean new file mode 100644 index 0000000000000..c236ef8443cf2 --- /dev/null +++ b/Mathlib/NumberTheory/ModularForms/QExpansion.lean @@ -0,0 +1,125 @@ +/- +Copyright (c) 2024 David Loeffler. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Loeffler +-/ +import Mathlib.Analysis.Complex.Periodic +import Mathlib.NumberTheory.ModularForms.Basic +import Mathlib.NumberTheory.ModularForms.Identities + +/-! +# q-expansions of modular forms + +We show that a modular form of level `Γ(n)` can be written as `τ ↦ F (𝕢 n τ)` where `F` is +analytic on the open unit disc, and `𝕢 n` is the parameter `τ ↦ exp (2 * I * π * τ / n)`. As an +application, we show that cusp forms decay exponentially to 0 as `im τ → ∞`. + +## TO DO: + +* generalise to handle arbitrary finite-index subgroups (not just `Γ(n)` for some `n`) +* define the `q`-expansion as a formal power series +-/ + +open ModularForm Complex Filter Asymptotics UpperHalfPlane Function + +open scoped Real Topology Manifold MatrixGroups CongruenceSubgroup + +noncomputable section + +variable {k : ℤ} {F : Type*} [FunLike F ℍ ℂ] {Γ : Subgroup SL(2, ℤ)} (n : ℕ) (f : F) + +local notation "I∞" => comap Complex.im atTop +local notation "𝕢" => Periodic.qParam + +theorem Function.Periodic.im_invQParam_pos_of_abs_lt_one + {h : ℝ} (hh : 0 < h) {q : ℂ} (hq : q.abs < 1) (hq_ne : q ≠ 0) : + 0 < im (Periodic.invQParam h q) := + im_invQParam .. ▸ mul_pos_of_neg_of_neg + (div_neg_of_neg_of_pos (neg_lt_zero.mpr hh) Real.two_pi_pos) + ((Real.log_neg_iff (Complex.abs.pos hq_ne)).mpr hq) + +namespace SlashInvariantFormClass + +theorem periodic_comp_ofComplex [SlashInvariantFormClass F Γ(n) k] : + Periodic (f ∘ ofComplex) n := by + intro w + by_cases hw : 0 < im w + · have : 0 < im (w + n) := by simp only [add_im, natCast_im, add_zero, hw] + simp only [comp_apply, ofComplex_apply_of_im_pos this, ofComplex_apply_of_im_pos hw] + convert SlashInvariantForm.vAdd_width_periodic n k 1 f ⟨w, hw⟩ using 2 + simp only [Int.cast_one, mul_one, UpperHalfPlane.ext_iff, coe_mk_subtype, coe_vadd, + ofReal_natCast, add_comm] + · have : im (w + n) ≤ 0 := by simpa only [add_im, natCast_im, add_zero, not_lt] using hw + simp only [comp_apply, ofComplex_apply_of_im_nonpos this, + ofComplex_apply_of_im_nonpos (not_lt.mp hw)] + +/-- +The analytic function `F` such that `f τ = F (exp (2 * π * I * τ / n))`, extended by a choice of +limit at `0`. +-/ +def cuspFunction : ℂ → ℂ := Function.Periodic.cuspFunction n (f ∘ ofComplex) + +theorem eq_cuspFunction [NeZero n] [SlashInvariantFormClass F Γ(n) k] (τ : ℍ) : + cuspFunction n f (𝕢 n τ) = f τ := by + simpa only [comp_apply, ofComplex_apply] + using (periodic_comp_ofComplex n f).eq_cuspFunction (NeZero.ne _) τ + +end SlashInvariantFormClass + +open SlashInvariantFormClass + +namespace ModularFormClass + +theorem differentiableAt_comp_ofComplex [ModularFormClass F Γ k] {z : ℂ} (hz : 0 < im z) : + DifferentiableAt ℂ (f ∘ ofComplex) z := + mdifferentiableAt_iff_differentiableAt.mp ((holo f _).comp z (mdifferentiableAt_ofComplex hz)) + +theorem bounded_at_infty_comp_ofComplex [ModularFormClass F Γ k] : + BoundedAtFilter I∞ (f ∘ ofComplex) := by + simpa only [SlashAction.slash_one, ModularForm.toSlashInvariantForm_coe] + using (ModularFormClass.bdd_at_infty f 1).comp_tendsto tendsto_comap_im_ofComplex + +theorem differentiableAt_cuspFunction [NeZero n] [ModularFormClass F Γ(n) k] + {q : ℂ} (hq : q.abs < 1) : + DifferentiableAt ℂ (cuspFunction n f) q := by + have npos : 0 < (n : ℝ) := mod_cast (Nat.pos_iff_ne_zero.mpr (NeZero.ne _)) + rcases eq_or_ne q 0 with rfl | hq' + · exact (periodic_comp_ofComplex n f).differentiableAt_cuspFunction_zero npos + (eventually_of_mem (preimage_mem_comap (Ioi_mem_atTop 0)) + (fun _ ↦ differentiableAt_comp_ofComplex f)) + (bounded_at_infty_comp_ofComplex f) + · exact Periodic.qParam_right_inv npos.ne' hq' ▸ + (periodic_comp_ofComplex n f).differentiableAt_cuspFunction npos.ne' + <| differentiableAt_comp_ofComplex _ <| Periodic.im_invQParam_pos_of_abs_lt_one npos hq hq' + +lemma analyticAt_cuspFunction_zero [NeZero n] [ModularFormClass F Γ(n) k] : + AnalyticAt ℂ (cuspFunction n f) 0 := + DifferentiableOn.analyticAt + (fun q hq ↦ (differentiableAt_cuspFunction _ _ hq).differentiableWithinAt) + (by simpa only [ball_zero_eq] using Metric.ball_mem_nhds (0 : ℂ) zero_lt_one) + +end ModularFormClass + +open ModularFormClass + +namespace CuspFormClass + +theorem zero_at_infty_comp_ofComplex [CuspFormClass F Γ k] : ZeroAtFilter I∞ (f ∘ ofComplex) := by + simpa only [SlashAction.slash_one, toSlashInvariantForm_coe] + using (zero_at_infty f 1).comp tendsto_comap_im_ofComplex + +theorem cuspFunction_apply_zero [NeZero n] [CuspFormClass F Γ(n) k] : + cuspFunction n f 0 = 0 := + Periodic.cuspFunction_zero_of_zero_at_inf (mod_cast (Nat.pos_iff_ne_zero.mpr (NeZero.ne _))) + (zero_at_infty_comp_ofComplex f) + +theorem exp_decay_atImInfty [NeZero n] [CuspFormClass F Γ(n) k] : + f =O[atImInfty] fun τ ↦ Real.exp (-2 * π * τ.im / n) := by + simpa only [neg_mul, comp_def, ofComplex_apply, coe_im] using + ((periodic_comp_ofComplex n f).exp_decay_of_zero_at_inf + (mod_cast (Nat.pos_iff_ne_zero.mpr (NeZero.ne _))) + (eventually_of_mem (preimage_mem_comap (Ioi_mem_atTop 0)) + fun _ ↦ differentiableAt_comp_ofComplex f) + (zero_at_infty_comp_ofComplex f)).comp_tendsto tendsto_coe_atImInfty + +end CuspFormClass From a0cd3d6f9e5bf54add78c47510188bfe2dc782bf Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Sat, 16 Nov 2024 22:07:58 +0000 Subject: [PATCH 019/829] chore(Rat/Cardinal): reduce imports and create new file (#18584) Co-authored-by: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/Data/Rat/Cardinal.lean | 22 ++++++++++++++++++++ Mathlib/Data/Rat/Denumerable.lean | 30 ++++++++-------------------- Mathlib/Data/Real/Cardinality.lean | 2 +- Mathlib/FieldTheory/Cardinality.lean | 2 +- Mathlib/ModelTheory/Order.lean | 3 ++- 6 files changed, 35 insertions(+), 25 deletions(-) create mode 100644 Mathlib/Data/Rat/Cardinal.lean diff --git a/Mathlib.lean b/Mathlib.lean index ef683fa1d3328..4902e66c634db 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2729,6 +2729,7 @@ import Mathlib.Data.QPF.Multivariate.Constructions.Sigma import Mathlib.Data.QPF.Univariate.Basic import Mathlib.Data.Quot import Mathlib.Data.Rat.BigOperators +import Mathlib.Data.Rat.Cardinal import Mathlib.Data.Rat.Cast.CharZero import Mathlib.Data.Rat.Cast.Defs import Mathlib.Data.Rat.Cast.Lemmas diff --git a/Mathlib/Data/Rat/Cardinal.lean b/Mathlib/Data/Rat/Cardinal.lean new file mode 100644 index 0000000000000..36332a7d1d4b5 --- /dev/null +++ b/Mathlib/Data/Rat/Cardinal.lean @@ -0,0 +1,22 @@ +/- +Copyright (c) 2019 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes, Floris Van Doorn +-/ +import Mathlib.Algebra.CharZero.Infinite +import Mathlib.Algebra.Ring.Rat +import Mathlib.Data.Rat.Encodable +import Mathlib.SetTheory.Cardinal.Basic + +/-! +# Cardinality of ℚ + +This file proves that the Cardinality of ℚ is ℵ₀ +-/ + +assert_not_exists Module +assert_not_exists Field + +open Cardinal + +theorem Cardinal.mkRat : #ℚ = ℵ₀ := mk_eq_aleph0 ℚ diff --git a/Mathlib/Data/Rat/Denumerable.lean b/Mathlib/Data/Rat/Denumerable.lean index 0a4c26dc0c331..9a1494cd76969 100644 --- a/Mathlib/Data/Rat/Denumerable.lean +++ b/Mathlib/Data/Rat/Denumerable.lean @@ -3,13 +3,17 @@ Copyright (c) 2019 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import Mathlib.Algebra.Order.Ring.Rat -import Mathlib.SetTheory.Cardinal.Basic +import Mathlib.Algebra.CharZero.Infinite +import Mathlib.Algebra.Ring.Rat +import Mathlib.Data.Rat.Encodable +import Mathlib.Logic.Denumerable /-! # Denumerability of ℚ -This file proves that ℚ is infinite, denumerable, and deduces that it has cardinality `omega`. +This file proves that ℚ is denumerable. + +The fact that ℚ has cardinality ℵ₀ is proved in `Data.Rat.Cardinal` -/ assert_not_exists Module @@ -19,25 +23,7 @@ namespace Rat open Denumerable -instance : Infinite ℚ := - Infinite.of_injective ((↑) : ℕ → ℚ) Nat.cast_injective - -private def denumerable_aux : ℚ ≃ { x : ℤ × ℕ // 0 < x.2 ∧ x.1.natAbs.Coprime x.2 } where - toFun x := ⟨⟨x.1, x.2⟩, Nat.pos_of_ne_zero x.3, x.4⟩ - invFun x := ⟨x.1.1, x.1.2, ne_zero_of_lt x.2.1, x.2.2⟩ - left_inv := fun ⟨_, _, _, _⟩ => rfl - right_inv := fun ⟨⟨_, _⟩, _, _⟩ => rfl - /-- **Denumerability of the Rational Numbers** -/ -instance instDenumerable : Denumerable ℚ := by - let T := { x : ℤ × ℕ // 0 < x.2 ∧ x.1.natAbs.Coprime x.2 } - letI : Infinite T := Infinite.of_injective _ denumerable_aux.injective - letI : Encodable T := Subtype.encodable - letI : Denumerable T := ofEncodableOfInfinite T - exact Denumerable.ofEquiv T denumerable_aux +instance instDenumerable : Denumerable ℚ := ofEncodableOfInfinite ℚ end Rat - -open Cardinal - -theorem Cardinal.mkRat : #ℚ = ℵ₀ := by simp only [mk_eq_aleph0] diff --git a/Mathlib/Data/Real/Cardinality.lean b/Mathlib/Data/Real/Cardinality.lean index 0ea501663c7c3..b0a7e99a26fb2 100644 --- a/Mathlib/Data/Real/Cardinality.lean +++ b/Mathlib/Data/Real/Cardinality.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ import Mathlib.Analysis.SpecificLimits.Basic -import Mathlib.Data.Rat.Denumerable +import Mathlib.Data.Rat.Cardinal import Mathlib.Data.Set.Pointwise.Interval import Mathlib.SetTheory.Cardinal.Continuum diff --git a/Mathlib/FieldTheory/Cardinality.lean b/Mathlib/FieldTheory/Cardinality.lean index 996728315a835..dfc11344f0654 100644 --- a/Mathlib/FieldTheory/Cardinality.lean +++ b/Mathlib/FieldTheory/Cardinality.lean @@ -6,7 +6,7 @@ Authors: Eric Rodriguez import Mathlib.Algebra.Field.ULift import Mathlib.Algebra.MvPolynomial.Cardinal import Mathlib.Data.Nat.Factorization.PrimePow -import Mathlib.Data.Rat.Denumerable +import Mathlib.Data.Rat.Encodable import Mathlib.FieldTheory.Finite.GaloisField import Mathlib.RingTheory.Localization.Cardinality import Mathlib.SetTheory.Cardinal.Divisibility diff --git a/Mathlib/ModelTheory/Order.lean b/Mathlib/ModelTheory/Order.lean index 02b6eaab6e02e..3e83acae8d3fb 100644 --- a/Mathlib/ModelTheory/Order.lean +++ b/Mathlib/ModelTheory/Order.lean @@ -3,7 +3,8 @@ Copyright (c) 2022 Aaron Anderson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ -import Mathlib.Data.Rat.Denumerable +import Mathlib.Algebra.CharZero.Infinite +import Mathlib.Data.Rat.Encodable import Mathlib.ModelTheory.Complexity import Mathlib.ModelTheory.Fraisse import Mathlib.Order.CountableDenseLinearOrder From 08fa0e9d6aedc38486d2c6f5f885c9d4b78c6743 Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Sat, 16 Nov 2024 22:27:17 +0000 Subject: [PATCH 020/829] chore(Algebra/Category): remove erws (#19132) --- Mathlib/Algebra/Category/Ring/Colimits.lean | 30 +++++++-------------- 1 file changed, 10 insertions(+), 20 deletions(-) diff --git a/Mathlib/Algebra/Category/Ring/Colimits.lean b/Mathlib/Algebra/Category/Ring/Colimits.lean index 626a251018780..140d787c20ebf 100644 --- a/Mathlib/Algebra/Category/Ring/Colimits.lean +++ b/Mathlib/Algebra/Category/Ring/Colimits.lean @@ -280,18 +280,13 @@ def colimitIsColimit : IsColimit (colimitCocone F) where refine Quot.inductionOn x ?_ intro x induction x with - | zero => erw [quot_zero, map_zero (f := m), (descMorphism F s).map_zero] - | one => erw [quot_one, map_one (f := m), (descMorphism F s).map_one] - -- extra rfl with https://github.com/leanprover/lean4/pull/2644 - | neg x ih => erw [quot_neg, map_neg (f := m), (descMorphism F s).map_neg, ih]; rfl + | zero => simp + | one => simp + | neg x ih => simp [ih] | of j x => exact congr_fun (congr_arg (fun f : F.obj j ⟶ s.pt => (f : F.obj j → s.pt)) (w j)) x - | add x y ih_x ih_y => - -- extra rfl with https://github.com/leanprover/lean4/pull/2644 - erw [quot_add, map_add (f := m), (descMorphism F s).map_add, ih_x, ih_y]; rfl - | mul x y ih_x ih_y => - -- extra rfl with https://github.com/leanprover/lean4/pull/2644 - erw [quot_mul, map_mul (f := m), (descMorphism F s).map_mul, ih_x, ih_y]; rfl + | add x y ih_x ih_y => simp [ih_x, ih_y] + | mul x y ih_x ih_y => simp [ih_x, ih_y] instance hasColimits_ringCat : HasColimits RingCat where has_colimits_of_shape _ _ := @@ -589,18 +584,13 @@ def colimitIsColimit : IsColimit (colimitCocone F) where refine Quot.inductionOn x ?_ intro x induction x with - | zero => erw [quot_zero, map_zero (f := m), (descMorphism F s).map_zero] - | one => erw [quot_one, map_one (f := m), (descMorphism F s).map_one] - -- extra rfl with https://github.com/leanprover/lean4/pull/2644 - | neg x ih => erw [quot_neg, map_neg (f := m), (descMorphism F s).map_neg, ih]; rfl + | zero => simp + | one => simp + | neg x ih => simp [ih] | of j x => exact congr_fun (congr_arg (fun f : F.obj j ⟶ s.pt => (f : F.obj j → s.pt)) (w j)) x - | add x y ih_x ih_y => - -- extra rfl with https://github.com/leanprover/lean4/pull/2644 - erw [quot_add, map_add (f := m), (descMorphism F s).map_add, ih_x, ih_y]; rfl - | mul x y ih_x ih_y => - -- extra rfl with https://github.com/leanprover/lean4/pull/2644 - erw [quot_mul, map_mul (f := m), (descMorphism F s).map_mul, ih_x, ih_y]; rfl + | add x y ih_x ih_y => simp [ih_x, ih_y] + | mul x y ih_x ih_y => simp [ih_x, ih_y] instance hasColimits_commRingCat : HasColimits CommRingCat where has_colimits_of_shape _ _ := From fe4de19cc7609bc70565b7a712202dc78ee54336 Mon Sep 17 00:00:00 2001 From: leanprover-community-bot-assistant Date: Sun, 17 Nov 2024 00:16:13 +0000 Subject: [PATCH 021/829] chore(scripts): update nolints.json (#19137) I am happy to remove some nolints for you! --- scripts/nolints.json | 3 --- 1 file changed, 3 deletions(-) diff --git a/scripts/nolints.json b/scripts/nolints.json index 10e10847b75d3..a64eacb27c056 100644 --- a/scripts/nolints.json +++ b/scripts/nolints.json @@ -149,9 +149,6 @@ ["docBlame", "FirstOrder.«term∀'_»"], ["docBlame", "FirstOrder.«term∃'_»"], ["docBlame", "FirstOrder.«term∼_»"], - ["docBlame", "Function.app"], - ["docBlame", "Function.compLeft"], - ["docBlame", "Function.compRight"], ["docBlame", "Function.swap"], ["docBlame", "Function.«term_∘'_»"], ["docBlame", "HSpace.e"], From e3327a6356ba7550fa37b50f16e934a4d59dc6af Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 17 Nov 2024 00:44:22 +0000 Subject: [PATCH 022/829] feat: a function with finite domain is bounded (#19134) This is particularly useful for anonymous dot notation From LeanAPAP --- Mathlib/Order/LiminfLimsup.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/Mathlib/Order/LiminfLimsup.lean b/Mathlib/Order/LiminfLimsup.lean index b9c9d64334b5e..25404163646dc 100644 --- a/Mathlib/Order/LiminfLimsup.lean +++ b/Mathlib/Order/LiminfLimsup.lean @@ -154,6 +154,14 @@ lemma _root_.BddBelow.isBoundedUnder (hs : s ∈ f) (hu : BddBelow (u '' s)) : lemma _root_.BddBelow.isBoundedUnder_of_range (hu : BddBelow (Set.range u)) : f.IsBoundedUnder (· ≥ ·) u := BddBelow.isBoundedUnder (s := univ) f.univ_mem (by simpa) +lemma IsBoundedUnder.le_of_finite [Nonempty α] [IsDirected α (· ≤ ·)] [Finite β] + {f : Filter β} {u : β → α} : IsBoundedUnder (· ≤ ·) f u := + (Set.toFinite _).bddAbove.isBoundedUnder_of_range + +lemma IsBoundedUnder.ge_of_finite [Nonempty α] [IsDirected α (· ≥ ·)] [Finite β] + {f : Filter β} {u : β → α} : IsBoundedUnder (· ≥ ·) f u := + (Set.toFinite _).bddBelow.isBoundedUnder_of_range + end Preorder theorem _root_.Monotone.isBoundedUnder_le_comp [Preorder α] [Preorder β] {l : Filter γ} {u : γ → α} From 59e75948eb18b0946b93c99d666844ca885675c6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sun, 17 Nov 2024 07:50:18 +0000 Subject: [PATCH 023/829] chore(SetTheory/Cardinal/Cofinality): golf `lift_cof` (#18047) We also drop some redundant TC assumptions in `RelIso.cof_eq_lift` and `RelIso.cof_eq`. --- Mathlib/SetTheory/Cardinal/Cofinality.lean | 53 +++++++++------------- 1 file changed, 22 insertions(+), 31 deletions(-) diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index 39c888270d868..d41d5f8e58c94 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -49,7 +49,7 @@ open scoped Ordinal universe u v w -variable {α : Type*} {r : α → α → Prop} +variable {α : Type u} {β : Type v} {r : α → α → Prop} {s : β → β → Prop} /-! ### Cofinality of orders -/ @@ -70,7 +70,7 @@ private theorem cof_nonempty (r : α → α → Prop) [IsRefl α r] : theorem cof_le (r : α → α → Prop) {S : Set α} (h : ∀ a, ∃ b ∈ S, r a b) : cof r ≤ #S := csInf_le' ⟨S, h, rfl⟩ -theorem le_cof {r : α → α → Prop} [IsRefl α r] (c : Cardinal) : +theorem le_cof [IsRefl α r] (c : Cardinal) : c ≤ cof r ↔ ∀ {S : Set α}, (∀ a, ∃ b ∈ S, r a b) → c ≤ #S := by rw [cof, le_csInf_iff'' (cof_nonempty r)] use fun H S h => H _ ⟨S, h, rfl⟩ @@ -79,8 +79,10 @@ theorem le_cof {r : α → α → Prop} [IsRefl α r] (c : Cardinal) : end Order -theorem RelIso.cof_le_lift {α : Type u} {β : Type v} {r : α → α → Prop} {s} [IsRefl β s] - (f : r ≃r s) : Cardinal.lift.{v} (Order.cof r) ≤ Cardinal.lift.{u} (Order.cof s) := by +namespace RelIso + +private theorem cof_le_lift [IsRefl β s] (f : r ≃r s) : + Cardinal.lift.{v} (Order.cof r) ≤ Cardinal.lift.{u} (Order.cof s) := by rw [Order.cof, Order.cof, lift_sInf, lift_sInf, le_csInf_iff'' ((Order.cof_nonempty s).image _)] rintro - ⟨-, ⟨u, H, rfl⟩, rfl⟩ apply csInf_le' @@ -89,17 +91,21 @@ theorem RelIso.cof_le_lift {α : Type u} {β : Type v} {r : α → α → Prop} refine ⟨f.symm b, mem_image_of_mem _ hb, f.map_rel_iff.1 ?_⟩ rwa [RelIso.apply_symm_apply] -theorem RelIso.cof_eq_lift {α : Type u} {β : Type v} {r s} [IsRefl α r] [IsRefl β s] (f : r ≃r s) : +theorem cof_eq_lift [IsRefl β s] (f : r ≃r s) : Cardinal.lift.{v} (Order.cof r) = Cardinal.lift.{u} (Order.cof s) := - (RelIso.cof_le_lift f).antisymm (RelIso.cof_le_lift f.symm) + have := f.toRelEmbedding.isRefl + (f.cof_le_lift).antisymm (f.symm.cof_le_lift) + +theorem cof_eq {α β : Type u} {r : α → α → Prop} {s} [IsRefl β s] (f : r ≃r s) : + Order.cof r = Order.cof s := + lift_inj.1 (f.cof_eq_lift) -theorem RelIso.cof_le {α β : Type u} {r : α → α → Prop} {s} [IsRefl β s] (f : r ≃r s) : +@[deprecated cof_eq (since := "2024-10-22")] +theorem cof_le {α β : Type u} {r : α → α → Prop} {s} [IsRefl β s] (f : r ≃r s) : Order.cof r ≤ Order.cof s := - lift_le.1 (RelIso.cof_le_lift f) + f.cof_eq.le -theorem RelIso.cof_eq {α β : Type u} {r s} [IsRefl α r] [IsRefl β s] (f : r ≃r s) : - Order.cof r = Order.cof s := - lift_inj.1 (RelIso.cof_eq_lift f) +end RelIso /-- Cofinality of a strict order `≺`. This is the smallest cardinality of a set `S : Set α` such that `∀ a, ∃ b ∈ S, ¬ b ≺ a`. -/ @@ -229,26 +235,11 @@ theorem cof_eq_sInf_lsub (o : Ordinal.{u}) : cof o = @[simp] theorem lift_cof (o) : Cardinal.lift.{u, v} (cof o) = cof (Ordinal.lift.{u, v} o) := by - refine inductionOn o ?_ - intro α r _ - apply le_antisymm - · refine le_cof_type.2 fun S H => ?_ - have : Cardinal.lift.{u, v} #(ULift.up ⁻¹' S) ≤ #(S : Type (max u v)) := by - rw [← Cardinal.lift_umax.{v, u}, ← Cardinal.lift_id'.{v, u} #S] - exact mk_preimage_of_injective_lift.{v, max u v} ULift.up S (ULift.up_injective.{v, u}) - refine (Cardinal.lift_le.2 <| cof_type_le ?_).trans this - exact fun a => - let ⟨⟨b⟩, bs, br⟩ := H ⟨a⟩ - ⟨b, bs, br⟩ - · rcases cof_eq r with ⟨S, H, e'⟩ - have : #(ULift.down.{u, v} ⁻¹' S) ≤ Cardinal.lift.{u, v} #S := - ⟨⟨fun ⟨⟨x⟩, h⟩ => ⟨⟨x, h⟩⟩, fun ⟨⟨x⟩, h₁⟩ ⟨⟨y⟩, h₂⟩ e => by - simp at e; congr⟩⟩ - rw [e'] at this - refine (cof_type_le ?_).trans this - exact fun ⟨a⟩ => - let ⟨b, bs, br⟩ := H a - ⟨⟨b⟩, bs, br⟩ + refine inductionOn o fun α r _ ↦ ?_ + rw [← type_uLift, cof_type, cof_type, ← Cardinal.lift_id'.{v, u} (Order.cof _), + ← Cardinal.lift_umax] + apply RelIso.cof_eq_lift ⟨Equiv.ulift.symm, _⟩ + simp [swap] theorem cof_le_card (o) : cof o ≤ card o := by rw [cof_eq_sInf_lsub] From 69311c4f51b73b6f823a10adb04e70fc635f8f08 Mon Sep 17 00:00:00 2001 From: Yongle Hu Date: Sun, 17 Nov 2024 11:30:53 +0000 Subject: [PATCH 024/829] chore(NumberTheory/RamificationInertia): change the variables in `inertiaDeg_algebra_tower` to explicit variables (#19138) Change the variables in [Ideal.inertiaDeg_algebra_tower](https://leanprover-community.github.io/mathlib4_docs/Mathlib/NumberTheory/RamificationInertia.html#Ideal.inertiaDeg_algebra_tower) to explicit variables, since no variables can be inferred. Co-authored-by: Hu Yongle --- Mathlib/NumberTheory/RamificationInertia.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/RamificationInertia.lean b/Mathlib/NumberTheory/RamificationInertia.lean index 976ece9c00b57..92cb53f8ad16f 100644 --- a/Mathlib/NumberTheory/RamificationInertia.lean +++ b/Mathlib/NumberTheory/RamificationInertia.lean @@ -849,7 +849,7 @@ theorem ramificationIdx_algebra_tower [IsDedekindDomain S] [IsDedekindDomain T] /-- Let `T / S / R` be a tower of algebras, `p, P, I` be ideals in `R, S, T`, respectively, and `p` and `P` are maximal. If `p = P ∩ S` and `P = I ∩ S`, then `f (I | p) = f (P | p) * f (I | P)`. -/ -theorem inertiaDeg_algebra_tower {p : Ideal R} {P : Ideal S} {I : Ideal T} [p.IsMaximal] +theorem inertiaDeg_algebra_tower (p : Ideal R) (P : Ideal S) (I : Ideal T) [p.IsMaximal] [P.IsMaximal] [P.LiesOver p] [I.LiesOver P] : inertiaDeg (algebraMap R T) p I = inertiaDeg (algebraMap R S) p P * inertiaDeg (algebraMap S T) P I := by rw [IsScalarTower.algebraMap_eq R S T] From 660c520d38232e635a37e679c44311add034df4c Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 17 Nov 2024 14:25:06 +0000 Subject: [PATCH 025/829] feat(KrullDimension): krullDim lemmas (#19111) Part of #15524. (There will be more interesting krullDim lemma once we add the full coheight api) From the Carleson project. --- Mathlib/Order/KrullDimension.lean | 77 +++++++++++++++++++++++++------ 1 file changed, 64 insertions(+), 13 deletions(-) diff --git a/Mathlib/Order/KrullDimension.lean b/Mathlib/Order/KrullDimension.lean index da2a6a15b28c5..bd5570bd4816a 100644 --- a/Mathlib/Order/KrullDimension.lean +++ b/Mathlib/Order/KrullDimension.lean @@ -411,22 +411,73 @@ lemma krullDim_eq_of_orderIso (f : α ≃o β) : krullDim α = krullDim β := /-- The Krull dimension is the supremum of the elements' heights. + +This version of the lemma assumes that `α` is nonempty. In this case, the coercion from `ℕ∞` to +`WithBot ℕ∞` is on the outside fo the right-hand side, which is usually more convenient. + +If `α` were empty, then `krullDim α = ⊥`. See `krullDim_eq_iSup_height` for the more general +version, with the coercion under the supremum. +-/ +lemma krullDim_eq_iSup_height_of_nonempty [Nonempty α] : krullDim α = ↑(⨆ (a : α), height a) := by + apply le_antisymm + · apply iSup_le + intro p + suffices p.length ≤ ⨆ (a : α), height a by + exact (WithBot.unbot'_le_iff fun _ => this).mp this + apply le_iSup_of_le p.last (length_le_height_last (p := p)) + · rw [krullDim_eq_iSup_length] + simp only [WithBot.coe_le_coe, iSup_le_iff] + intro x + exact height_le fun p _ ↦ le_iSup_of_le p le_rfl + +/-- +The Krull dimension is the supremum of the elements' coheights. + +This version of the lemma assumes that `α` is nonempty. In this case, the coercion from `ℕ∞` to +`WithBot ℕ∞` is on the outside of the right-hand side, which is usually more convenient. + +If `α` were empty, then `krullDim α = ⊥`. See `krullDim_eq_iSup_coheight` for the more general +version, with the coercion under the supremum. +-/ +lemma krullDim_eq_iSup_coheight_of_nonempty [Nonempty α] : + krullDim α = ↑(⨆ (a : α), coheight a) := by + rw [← krullDim_orderDual] + exact krullDim_eq_iSup_height_of_nonempty (α := αᵒᵈ) + +/-- +The Krull dimension is the supremum of the elements' heights. + +If `α` is `Nonempty`, then `krullDim_eq_iSup_height_of_nonempty`, with the coercion from +`ℕ∞` to `WithBot ℕ∞` outside the supremum, can be more convenient. -/ lemma krullDim_eq_iSup_height : krullDim α = ⨆ (a : α), ↑(height a) := by cases isEmpty_or_nonempty α with - | inl h => simp [krullDim_eq_bot_of_isEmpty] - | inr h => - rw [← WithBot.coe_iSup (OrderTop.bddAbove _)] - apply le_antisymm - · apply iSup_le - intro p - suffices p.length ≤ ⨆ (a : α), height a by - exact (WithBot.unbot'_le_iff fun _ => this).mp this - apply le_iSup_of_le p.last (length_le_height_last (p := p)) - · rw [krullDim_eq_iSup_length] - simp only [WithBot.coe_le_coe, iSup_le_iff] - intro x - exact height_le fun p _ ↦ le_iSup_of_le p le_rfl + | inl h => rw [krullDim_eq_bot_of_isEmpty, ciSup_of_empty] + | inr h => rw [krullDim_eq_iSup_height_of_nonempty, WithBot.coe_iSup (OrderTop.bddAbove _)] + +/-- +The Krull dimension is the supremum of the elements' coheights. + +If `α` is `Nonempty`, then `krullDim_eq_iSup_coheight_of_nonempty`, with the coercion from +`ℕ∞` to `WithBot ℕ∞` outside the supremum, can be more convenient. +-/ +lemma krullDim_eq_iSup_coheight : krullDim α = ⨆ (a : α), ↑(coheight a) := by + cases isEmpty_or_nonempty α with + | inl h => rw [krullDim_eq_bot_of_isEmpty, ciSup_of_empty] + | inr h => rw [krullDim_eq_iSup_coheight_of_nonempty, WithBot.coe_iSup (OrderTop.bddAbove _)] + +@[simp] -- not as useful as a simp lemma as it looks, due to the coe on the left +lemma height_top_eq_krullDim [OrderTop α] : height (⊤ : α) = krullDim α := by + rw [krullDim_eq_iSup_length] + simp only [WithBot.coe_inj] + apply le_antisymm + · exact height_le fun p _ ↦ le_iSup_of_le p le_rfl + · exact iSup_le fun _ => length_le_height le_top + +@[simp] -- not as useful as a simp lemma as it looks, due to the coe on the left +lemma coheight_bot_eq_krullDim [OrderBot α] : coheight (⊥ : α) = krullDim α := by + rw [← krullDim_orderDual] + exact height_top_eq_krullDim (α := αᵒᵈ) end krullDim From c0c8f97c406378fda228e9fba5a735b31f1e438d Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 17 Nov 2024 15:00:45 +0000 Subject: [PATCH 026/829] feat: add `measure_diff_symm` (#19049) --- Mathlib/MeasureTheory/Measure/NullMeasurable.lean | 10 +++++++++- 1 file changed, 9 insertions(+), 1 deletion(-) diff --git a/Mathlib/MeasureTheory/Measure/NullMeasurable.lean b/Mathlib/MeasureTheory/Measure/NullMeasurable.lean index 36cb6a66965c3..adefe7bb1eecd 100644 --- a/Mathlib/MeasureTheory/Measure/NullMeasurable.lean +++ b/Mathlib/MeasureTheory/Measure/NullMeasurable.lean @@ -55,8 +55,8 @@ the output type. measurable, measure, null measurable, completion -/ - open Filter Set Encodable +open scoped ENNReal variable {ι α β γ : Type*} @@ -260,6 +260,14 @@ theorem measure_inter_add_diff₀ (s : Set α) (ht : NullMeasurableSet t μ) : _ = μ s' := congr_arg μ (inter_union_diff _ _) _ = μ s := hs' +/-- If `s` and `t` are null measurable sets of equal measure +and their intersection has finite measure, +then `s \ t` and `t \ s` have equal measures too. -/ +theorem measure_diff_symm (hs : NullMeasurableSet s μ) (ht : NullMeasurableSet t μ) + (h : μ s = μ t) (hfin : μ (s ∩ t) ≠ ∞) : μ (s \ t) = μ (t \ s) := by + rw [← ENNReal.add_right_inj hfin, measure_inter_add_diff₀ _ ht, inter_comm, + measure_inter_add_diff₀ _ hs, h] + theorem measure_union_add_inter₀ (s : Set α) (ht : NullMeasurableSet t μ) : μ (s ∪ t) + μ (s ∩ t) = μ s + μ t := by rw [← measure_inter_add_diff₀ (s ∪ t) ht, union_inter_cancel_right, union_diff_right, ← From 32240b2510fc82a9b2ee4c3da4f79aff974a427a Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Sun, 17 Nov 2024 15:35:07 +0000 Subject: [PATCH 027/829] feat: the bicentralizer of a commutative set is commutative (#18700) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This shows that the bicentralizer (a.k.a. bicommutant) of a commutative set is commutative. Moreover, since `s ⊆ s.centralizer.centralizer`, if `s` is a commutative set, then `closure s ≤ centralizer (centralizer s)` for various subobject closures. Consequently, we obtain simplified proofs that if `s` is commutative, then so is `closure s`. Co-authored-by: ADedecker --- Mathlib/Algebra/Group/Center.lean | 9 ++++ Mathlib/Algebra/Group/Subgroup/Basic.lean | 21 -------- .../Algebra/Group/Submonoid/Membership.lean | 14 ----- Mathlib/Algebra/Ring/Subring/Basic.lean | 24 +++------ Mathlib/Algebra/Ring/Subsemiring/Basic.lean | 24 ++++----- Mathlib/Algebra/Star/Center.lean | 13 +++++ Mathlib/Algebra/Star/Subalgebra.lean | 53 +++++++++++-------- Mathlib/GroupTheory/Subgroup/Centralizer.lean | 17 ++++++ .../GroupTheory/Submonoid/Centralizer.lean | 17 ++++++ .../GroupTheory/Subsemigroup/Centralizer.lean | 17 ++++++ Mathlib/RingTheory/Adjoin/Basic.lean | 26 ++++----- .../NonUnitalSubsemiring/Basic.lean | 13 +++++ 12 files changed, 143 insertions(+), 105 deletions(-) diff --git a/Mathlib/Algebra/Group/Center.lean b/Mathlib/Algebra/Group/Center.lean index d5fd79ac44d8c..3af32138e4920 100644 --- a/Mathlib/Algebra/Group/Center.lean +++ b/Mathlib/Algebra/Group/Center.lean @@ -138,6 +138,10 @@ theorem mul_mem_center {z₁ z₂ : M} (hz₁ : z₁ ∈ Set.center M) (hz₂ : lemma center_subset_centralizer (S : Set M) : Set.center M ⊆ S.centralizer := fun _ hx m _ ↦ (hx.comm m).symm +@[to_additive addCentralizer_union] +lemma centralizer_union : centralizer (S ∪ T) = centralizer S ∩ centralizer T := by + simp [centralizer, or_imp, forall_and, setOf_and] + @[to_additive (attr := gcongr) addCentralizer_subset] lemma centralizer_subset (h : S ⊆ T) : centralizer T ⊆ centralizer S := fun _ ht s hs ↦ ht s (h hs) @@ -161,6 +165,11 @@ lemma centralizer_centralizer_centralizer (S : Set M) : instance decidableMemCentralizer [∀ a : M, Decidable <| ∀ b ∈ S, b * a = a * b] : DecidablePred (· ∈ centralizer S) := fun _ ↦ decidable_of_iff' _ mem_centralizer_iff +@[to_additive addCentralizer_addCentralizer_comm_of_comm] +lemma centralizer_centralizer_comm_of_comm (h_comm : ∀ x ∈ S, ∀ y ∈ S, x * y = y * x) : + ∀ x ∈ S.centralizer.centralizer, ∀ y ∈ S.centralizer.centralizer, x * y = y * x := + fun _ h₁ _ h₂ ↦ h₂ _ fun _ h₃ ↦ h₁ _ fun _ h₄ ↦ h_comm _ h₄ _ h₃ + end Mul section Semigroup diff --git a/Mathlib/Algebra/Group/Subgroup/Basic.lean b/Mathlib/Algebra/Group/Subgroup/Basic.lean index 866249716b622..6fc988e05de4b 100644 --- a/Mathlib/Algebra/Group/Subgroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subgroup/Basic.lean @@ -436,27 +436,6 @@ theorem closure_closure_coe_preimage {k : Set G} : closure (((↑) : closure k closure_induction (fun _ h ↦ subset_closure h) (one_mem _) (fun _ _ _ _ ↦ mul_mem) (fun _ _ ↦ inv_mem) hx' -/-- If all the elements of a set `s` commute, then `closure s` is a commutative group. -/ -@[to_additive - "If all the elements of a set `s` commute, then `closure s` is an additive - commutative group."] -def closureCommGroupOfComm {k : Set G} (hcomm : ∀ x ∈ k, ∀ y ∈ k, x * y = y * x) : - CommGroup (closure k) := - { (closure k).toGroup with - mul_comm := fun ⟨x, hx⟩ ⟨y, hy⟩ => by - ext - simp only [Subgroup.coe_mul] - induction hx, hy using closure_induction₂ with - | mem x y hx hy => exact hcomm x hx y hy - | one_left x _ => exact Commute.one_left x - | one_right x _ => exact Commute.one_right x - | mul_left _ _ _ _ _ _ h₁ h₂ => exact Commute.mul_left h₁ h₂ - | mul_right _ _ _ _ _ _ h₁ h₂ => exact Commute.mul_right h₁ h₂ - | inv_left _ _ _ _ h => -- `Commute.inv_left` is not imported - rw [inv_mul_eq_iff_eq_mul, ← mul_assoc, h, mul_assoc, mul_inv_cancel, mul_one] - | inv_right _ _ _ _ h => - rw [mul_inv_eq_iff_eq_mul, mul_assoc, h, ← mul_assoc, inv_mul_cancel, one_mul] } - variable (G) /-- `closure` forms a Galois insertion with the coercion to set. -/ diff --git a/Mathlib/Algebra/Group/Submonoid/Membership.lean b/Mathlib/Algebra/Group/Submonoid/Membership.lean index af500a66d1eff..cc9d5be5714ae 100644 --- a/Mathlib/Algebra/Group/Submonoid/Membership.lean +++ b/Mathlib/Algebra/Group/Submonoid/Membership.lean @@ -494,20 +494,6 @@ theorem map_powers {N : Type*} {F : Type*} [Monoid N] [FunLike F M N] [MonoidHom (powers m).map f = powers (f m) := by simp only [powers_eq_closure, map_mclosure f, Set.image_singleton] -/-- If all the elements of a set `s` commute, then `closure s` is a commutative monoid. -/ -@[to_additive - "If all the elements of a set `s` commute, then `closure s` forms an additive - commutative monoid."] -def closureCommMonoidOfComm {s : Set M} (hcomm : ∀ a ∈ s, ∀ b ∈ s, a * b = b * a) : - CommMonoid (closure s) := - { (closure s).toMonoid with - mul_comm := fun x y => by - ext - simp only [Submonoid.coe_mul] - exact closure_induction₂ (fun _ _ h₁ h₂ ↦ hcomm _ h₁ _ h₂) (fun _ _ ↦ Commute.one_left _) - (fun _ _ ↦ Commute.one_right _) (fun _ _ _ _ _ _ ↦ Commute.mul_left) - (fun _ _ _ _ _ _ ↦ Commute.mul_right) x.prop y.prop } - end Submonoid @[to_additive] diff --git a/Mathlib/Algebra/Ring/Subring/Basic.lean b/Mathlib/Algebra/Ring/Subring/Basic.lean index 017ce96f7a3e8..791ac77d64f46 100644 --- a/Mathlib/Algebra/Ring/Subring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subring/Basic.lean @@ -548,25 +548,17 @@ theorem mem_closure_iff {s : Set R} {x} : | mul _ _ _ _ h₁ h₂ => exact add_mem h₁ h₂ | inv _ _ h => exact neg_mem h⟩ +lemma closure_le_centralizer_centralizer (s : Set R) : + closure s ≤ centralizer (centralizer s) := + closure_le.mpr Set.subset_centralizer_centralizer + /-- If all elements of `s : Set A` commute pairwise, then `closure s` is a commutative ring. -/ -def closureCommRingOfComm {s : Set R} (hcomm : ∀ a ∈ s, ∀ b ∈ s, a * b = b * a) : +abbrev closureCommRingOfComm {s : Set R} (hcomm : ∀ x ∈ s, ∀ y ∈ s, x * y = y * x) : CommRing (closure s) := { (closure s).toRing with - mul_comm := fun ⟨x, hx⟩ ⟨y, hy⟩ => by - ext - simp only [MulMemClass.mk_mul_mk] - induction hx, hy using closure_induction₂ with - | mem_mem x y hx hy => exact hcomm x hx y hy - | zero_left x _ => exact Commute.zero_left x - | zero_right x _ => exact Commute.zero_right x - | one_left x _ => exact Commute.one_left x - | one_right x _ => exact Commute.one_right x - | mul_left _ _ _ _ _ _ h₁ h₂ => exact Commute.mul_left h₁ h₂ - | mul_right _ _ _ _ _ _ h₁ h₂ => exact Commute.mul_right h₁ h₂ - | add_left _ _ _ _ _ _ h₁ h₂ => exact Commute.add_left h₁ h₂ - | add_right _ _ _ _ _ _ h₁ h₂ => exact Commute.add_right h₁ h₂ - | neg_left _ _ _ _ h => exact Commute.neg_left h - | neg_right _ _ _ _ h => exact Commute.neg_right h } + mul_comm := fun ⟨_, h₁⟩ ⟨_, h₂⟩ ↦ + have := closure_le_centralizer_centralizer s + Subtype.ext <| Set.centralizer_centralizer_comm_of_comm hcomm _ (this h₁) _ (this h₂) } theorem exists_list_of_mem_closure {s : Set R} {x : R} (hx : x ∈ closure s) : ∃ L : List (List R), (∀ t ∈ L, ∀ y ∈ t, y ∈ s ∨ y = (-1 : R)) ∧ (L.map List.prod).sum = x := by diff --git a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean index c98fc5fd857a1..d955416892c36 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean @@ -886,23 +886,17 @@ instance center.smulCommClass_left : SMulCommClass (center R') R' R' := instance center.smulCommClass_right : SMulCommClass R' (center R') R' := Submonoid.center.smulCommClass_right -/-- If all the elements of a set `s` commute, then `closure s` is a commutative monoid. -/ -def closureCommSemiringOfComm {s : Set R'} (hcomm : ∀ a ∈ s, ∀ b ∈ s, a * b = b * a) : +lemma closure_le_centralizer_centralizer (s : Set R') : + closure s ≤ centralizer (centralizer s) := + closure_le.mpr Set.subset_centralizer_centralizer + +/-- If all the elements of a set `s` commute, then `closure s` is a commutative semiring. -/ +abbrev closureCommSemiringOfComm {s : Set R'} (hcomm : ∀ x ∈ s, ∀ y ∈ s, x * y = y * x) : CommSemiring (closure s) := { (closure s).toSemiring with - mul_comm := fun ⟨x, hx⟩ ⟨y, hy⟩ => by - ext - simp only [MulMemClass.mk_mul_mk] - induction hx, hy using closure_induction₂ with - | mem_mem x y hx hy => exact hcomm x hx y hy - | zero_left x _ => exact Commute.zero_left x - | zero_right x _ => exact Commute.zero_right x - | one_left x _ => exact Commute.one_left x - | one_right x _ => exact Commute.one_right x - | mul_left _ _ _ _ _ _ h₁ h₂ => exact Commute.mul_left h₁ h₂ - | mul_right _ _ _ _ _ _ h₁ h₂ => exact Commute.mul_right h₁ h₂ - | add_left _ _ _ _ _ _ h₁ h₂ => exact Commute.add_left h₁ h₂ - | add_right _ _ _ _ _ _ h₁ h₂ => exact Commute.add_right h₁ h₂ } + mul_comm := fun ⟨_, h₁⟩ ⟨_, h₂⟩ ↦ + have := closure_le_centralizer_centralizer s + Subtype.ext <| Set.centralizer_centralizer_comm_of_comm hcomm _ (this h₁) _ (this h₂) } end Subsemiring diff --git a/Mathlib/Algebra/Star/Center.lean b/Mathlib/Algebra/Star/Center.lean index b1ea404208648..c7390f8ad6704 100644 --- a/Mathlib/Algebra/Star/Center.lean +++ b/Mathlib/Algebra/Star/Center.lean @@ -33,6 +33,19 @@ theorem Set.star_mem_center (ha : a ∈ Set.center R) : star a ∈ Set.center R _ = b * star (a * star c) := by rw [star_mul, star_star] _ = b * (c * star a) := by rw [star_mul, star_star] +theorem Set.star_centralizer : star s.centralizer = (star s).centralizer := by + simp_rw [centralizer, ← commute_iff_eq] + conv_lhs => simp only [← star_preimage, preimage_setOf_eq, ← commute_star_comm] + conv_rhs => simp only [← image_star, forall_mem_image] + +theorem Set.union_star_self_comm (hcomm : ∀ x ∈ s, ∀ y ∈ s, y * x = x * y) + (hcomm_star : ∀ x ∈ s, ∀ y ∈ s, y * star x = star x * y) : + ∀ x ∈ s ∪ star s, ∀ y ∈ s ∪ star s, y * x = x * y := by + change s ∪ star s ⊆ (s ∪ star s).centralizer + simp_rw [centralizer_union, ← star_centralizer, union_subset_iff, subset_inter_iff, + star_subset_star, star_subset] + exact ⟨⟨hcomm, hcomm_star⟩, ⟨hcomm_star, hcomm⟩⟩ + theorem Set.star_mem_centralizer' (h : ∀ a : R, a ∈ s → star a ∈ s) (ha : a ∈ Set.centralizer s) : star a ∈ Set.centralizer s := fun y hy => by simpa using congr_arg star (ha _ (h _ hy)).symm diff --git a/Mathlib/Algebra/Star/Subalgebra.lean b/Mathlib/Algebra/Star/Subalgebra.lean index b2c1cc0ea0538..7df842838e470 100644 --- a/Mathlib/Algebra/Star/Subalgebra.lean +++ b/Mathlib/Algebra/Star/Subalgebra.lean @@ -265,16 +265,22 @@ def centralizer (s : Set A) : StarSubalgebra R A where theorem coe_centralizer (s : Set A) : (centralizer R s : Set A) = (s ∪ star s).centralizer := rfl -theorem mem_centralizer_iff {s : Set A} {z : A} : +open Set in +nonrec theorem mem_centralizer_iff {s : Set A} {z : A} : z ∈ centralizer R s ↔ ∀ g ∈ s, g * z = z * g ∧ star g * z = z * star g := by - show (∀ g ∈ s ∪ star s, g * z = z * g) ↔ ∀ g ∈ s, g * z = z * g ∧ star g * z = z * star g - simp only [Set.mem_union, or_imp, forall_and, and_congr_right_iff] - exact fun _ => - ⟨fun hz a ha => hz _ (Set.star_mem_star.mpr ha), fun hz a ha => star_star a ▸ hz _ ha⟩ + simp [← SetLike.mem_coe, centralizer_union, ← image_star, mem_centralizer_iff, forall_and] theorem centralizer_le (s t : Set A) (h : s ⊆ t) : centralizer R t ≤ centralizer R s := Set.centralizer_subset (Set.union_subset_union h <| Set.preimage_mono h) +theorem centralizer_toSubalgebra (s : Set A) : + (centralizer R s).toSubalgebra = Subalgebra.centralizer R (s ∪ star s):= + rfl + +theorem coe_centralizer_centralizer (s : Set A) : + (centralizer R (centralizer R s : Set A)) = (s ∪ star s).centralizer.centralizer := by + rw [coe_centralizer, StarMemClass.star_coe_eq, Set.union_self, coe_centralizer] + end Centralizer end StarSubalgebra @@ -503,28 +509,29 @@ theorem adjoin_induction_subtype {s : Set A} {p : adjoin R s → Prop} (a : adjo variable (R) +lemma adjoin_le_centralizer_centralizer (s : Set A) : + adjoin R s ≤ centralizer R (centralizer R s) := by + rw [← toSubalgebra_le_iff, centralizer_toSubalgebra, adjoin_toSubalgebra] + convert Algebra.adjoin_le_centralizer_centralizer R (s ∪ star s) + rw [StarMemClass.star_coe_eq] + simp + /-- If all elements of `s : Set A` commute pairwise and also commute pairwise with elements of `star s`, then `StarSubalgebra.adjoin R s` is commutative. See note [reducible non-instances]. -/ abbrev adjoinCommSemiringOfComm {s : Set A} - (hcomm : ∀ a : A, a ∈ s → ∀ b : A, b ∈ s → a * b = b * a) - (hcomm_star : ∀ a : A, a ∈ s → ∀ b : A, b ∈ s → a * star b = star b * a) : + (hcomm : ∀ a ∈ s, ∀ b ∈ s, a * b = b * a) + (hcomm_star : ∀ a ∈ s, ∀ b ∈ s, a * star b = star b * a) : CommSemiring (adjoin R s) := - { (adjoin R s).toSubalgebra.toSemiring with - mul_comm := by - rintro ⟨x, hx⟩ ⟨y, hy⟩ - ext - simp only [MulMemClass.mk_mul_mk] - rw [← mem_toSubalgebra, adjoin_toSubalgebra] at hx hy - letI : CommSemiring (Algebra.adjoin R (s ∪ star s)) := - Algebra.adjoinCommSemiringOfComm R - (by - intro a ha b hb - cases' ha with ha ha <;> cases' hb with hb hb - · exact hcomm _ ha _ hb - · exact star_star b ▸ hcomm_star _ ha _ hb - · exact star_star a ▸ (hcomm_star _ hb _ ha).symm - · simpa only [star_mul, star_star] using congr_arg star (hcomm _ hb _ ha)) - exact congr_arg Subtype.val (mul_comm (⟨x, hx⟩ : Algebra.adjoin R (s ∪ star s)) ⟨y, hy⟩) } + { (adjoin R s).toSemiring with + mul_comm := fun ⟨_, h₁⟩ ⟨_, h₂⟩ ↦ by + have hcomm : ∀ a ∈ s ∪ star s, ∀ b ∈ s ∪ star s, a * b = b * a := fun a ha b hb ↦ + Set.union_star_self_comm (fun _ ha _ hb ↦ hcomm _ hb _ ha) + (fun _ ha _ hb ↦ hcomm_star _ hb _ ha) b hb a ha + have := adjoin_le_centralizer_centralizer R s + apply this at h₁ + apply this at h₂ + rw [← SetLike.mem_coe, coe_centralizer_centralizer] at h₁ h₂ + exact Subtype.ext <| Set.centralizer_centralizer_comm_of_comm hcomm _ h₁ _ h₂ } /-- If all elements of `s : Set A` commute pairwise and also commute pairwise with elements of `star s`, then `StarSubalgebra.adjoin R s` is commutative. See note [reducible non-instances]. -/ diff --git a/Mathlib/GroupTheory/Subgroup/Centralizer.lean b/Mathlib/GroupTheory/Subgroup/Centralizer.lean index 6b5554c08a111..664c66c389c5c 100644 --- a/Mathlib/GroupTheory/Subgroup/Centralizer.lean +++ b/Mathlib/GroupTheory/Subgroup/Centralizer.lean @@ -77,4 +77,21 @@ variable (H) theorem le_centralizer [h : H.IsCommutative] : H ≤ centralizer H := le_centralizer_iff_isCommutative.mpr h +variable {H} in +@[to_additive] +lemma closure_le_centralizer_centralizer (s : Set G) : + closure s ≤ centralizer (centralizer s) := + closure_le _ |>.mpr Set.subset_centralizer_centralizer + +/-- If all the elements of a set `s` commute, then `closure s` is a commutative group. -/ +@[to_additive + "If all the elements of a set `s` commute, then `closure s` is an additive + commutative group."] +abbrev closureCommGroupOfComm {k : Set G} (hcomm : ∀ x ∈ k, ∀ y ∈ k, x * y = y * x) : + CommGroup (closure k) := + { (closure k).toGroup with + mul_comm := fun ⟨_, h₁⟩ ⟨_, h₂⟩ ↦ + have := closure_le_centralizer_centralizer k + Subtype.ext <| Set.centralizer_centralizer_comm_of_comm hcomm _ (this h₁) _ (this h₂) } + end Subgroup diff --git a/Mathlib/GroupTheory/Submonoid/Centralizer.lean b/Mathlib/GroupTheory/Submonoid/Centralizer.lean index 3ad0ece4f0e49..711ebf5d1a86b 100644 --- a/Mathlib/GroupTheory/Submonoid/Centralizer.lean +++ b/Mathlib/GroupTheory/Submonoid/Centralizer.lean @@ -88,6 +88,23 @@ lemma centralizer_centralizer_centralizer {s : Set M} : apply SetLike.coe_injective simp only [coe_centralizer, Set.centralizer_centralizer_centralizer] +variable {M} in +@[to_additive] +lemma closure_le_centralizer_centralizer (s : Set M) : + closure s ≤ centralizer (centralizer s) := + closure_le.mpr Set.subset_centralizer_centralizer + +/-- If all the elements of a set `s` commute, then `closure s` is a commutative monoid. -/ +@[to_additive + "If all the elements of a set `s` commute, then `closure s` forms an additive + commutative monoid."] +abbrev closureCommMonoidOfComm {s : Set M} (hcomm : ∀ a ∈ s, ∀ b ∈ s, a * b = b * a) : + CommMonoid (closure s) := + { (closure s).toMonoid with + mul_comm := fun ⟨_, h₁⟩ ⟨_, h₂⟩ ↦ + have := closure_le_centralizer_centralizer s + Subtype.ext <| Set.centralizer_centralizer_comm_of_comm hcomm _ (this h₁) _ (this h₂) } + end end Submonoid diff --git a/Mathlib/GroupTheory/Subsemigroup/Centralizer.lean b/Mathlib/GroupTheory/Subsemigroup/Centralizer.lean index 3f5aa2075a30e..b68aa6375a360 100644 --- a/Mathlib/GroupTheory/Subsemigroup/Centralizer.lean +++ b/Mathlib/GroupTheory/Subsemigroup/Centralizer.lean @@ -67,6 +67,23 @@ variable (M) theorem centralizer_univ : centralizer Set.univ = center M := SetLike.ext' (Set.centralizer_univ M) +variable {M} in +@[to_additive] +lemma closure_le_centralizer_centralizer (s : Set M) : + closure s ≤ centralizer (centralizer s) := + closure_le.mpr Set.subset_centralizer_centralizer + +/-- If all the elements of a set `s` commute, then `closure s` is a commutative semigroup. -/ +@[to_additive + "If all the elements of a set `s` commute, then `closure s` forms an additive + commutative semigroup."] +abbrev closureCommSemigroupOfComm {s : Set M} (hcomm : ∀ a ∈ s, ∀ b ∈ s, a * b = b * a) : + CommSemigroup (closure s) := + { MulMemClass.toSemigroup (closure s) with + mul_comm := fun ⟨_, h₁⟩ ⟨_, h₂⟩ ↦ + have := closure_le_centralizer_centralizer s + Subtype.ext <| Set.centralizer_centralizer_comm_of_comm hcomm _ (this h₁) _ (this h₂) } + end end Subsemigroup diff --git a/Mathlib/RingTheory/Adjoin/Basic.lean b/Mathlib/RingTheory/Adjoin/Basic.lean index 61c356ad30636..aa670454a746e 100644 --- a/Mathlib/RingTheory/Adjoin/Basic.lean +++ b/Mathlib/RingTheory/Adjoin/Basic.lean @@ -255,23 +255,17 @@ theorem adjoin_inl_union_inr_eq_prod (s) (t) : replace Hb : ((0 : A), b) ∈ P := adjoin_mono Set.subset_union_right Hb simpa [P] using Subalgebra.add_mem _ Ha Hb -/-- If all elements of `s : Set A` commute pairwise, then `adjoin R s` is a commutative -semiring. -/ -def adjoinCommSemiringOfComm {s : Set A} (hcomm : ∀ a ∈ s, ∀ b ∈ s, a * b = b * a) : +lemma adjoin_le_centralizer_centralizer (s : Set A) : + adjoin R s ≤ Subalgebra.centralizer R (Subalgebra.centralizer R s) := + adjoin_le Set.subset_centralizer_centralizer + +/-- If all elements of `s : Set A` commute pairwise, then `adjoin s` is a commutative semiring. -/ +abbrev adjoinCommSemiringOfComm {s : Set A} (hcomm : ∀ a ∈ s, ∀ b ∈ s, a * b = b * a) : CommSemiring (adjoin R s) := { (adjoin R s).toSemiring with - mul_comm := fun ⟨x, hx⟩ ⟨y, hy⟩ => by - ext - simp only [MulMemClass.mk_mul_mk] - induction hx, hy using adjoin_induction₂ with - | mem_mem x y hx hy => exact hcomm x hx y hy - | algebraMap_both r₁ r₂ => exact commutes r₁ <| algebraMap R A r₂ - | algebraMap_left r x _ => exact commutes r x - | algebraMap_right r x _ => exact commutes r x |>.symm - | mul_left _ _ _ _ _ _ h₁ h₂ => exact Commute.mul_left h₁ h₂ - | mul_right _ _ _ _ _ _ h₁ h₂ => exact Commute.mul_right h₁ h₂ - | add_left _ _ _ _ _ _ h₁ h₂ => exact Commute.add_left h₁ h₂ - | add_right _ _ _ _ _ _ h₁ h₂ => exact Commute.add_right h₁ h₂ } + mul_comm := fun ⟨_, h₁⟩ ⟨_, h₂⟩ ↦ + have := adjoin_le_centralizer_centralizer R s + Subtype.ext <| Set.centralizer_centralizer_comm_of_comm hcomm _ (this h₁) _ (this h₂) } variable {R} @@ -415,7 +409,7 @@ variable (R) /-- If all elements of `s : Set A` commute pairwise, then `adjoin R s` is a commutative ring. -/ -def adjoinCommRingOfComm {s : Set A} (hcomm : ∀ a ∈ s, ∀ b ∈ s, a * b = b * a) : +abbrev adjoinCommRingOfComm {s : Set A} (hcomm : ∀ a ∈ s, ∀ b ∈ s, a * b = b * a) : CommRing (adjoin R s) := { (adjoin R s).toRing, adjoinCommSemiringOfComm R hcomm with } diff --git a/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean b/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean index e008a205dc742..71ae33dcc8107 100644 --- a/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean +++ b/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean @@ -337,6 +337,19 @@ theorem closure_eq_of_le {s : Set R} {t : NonUnitalSubsemiring R} (h₁ : s ⊆ (h₂ : t ≤ closure s) : closure s = t := le_antisymm (closure_le.2 h₁) h₂ +lemma closure_le_centralizer_centralizer {R : Type*} [NonUnitalSemiring R] (s : Set R) : + closure s ≤ centralizer (centralizer s) := + closure_le.mpr Set.subset_centralizer_centralizer + +/-- If all the elements of a set `s` commute, then `closure s` is a non-unital commutative +semiring. -/ +abbrev closureNonUnitalCommSemiringOfComm {R : Type*} [NonUnitalSemiring R] {s : Set R} + (hcomm : ∀ x ∈ s, ∀ y ∈ s, x * y = y * x) : NonUnitalCommSemiring (closure s) := + { NonUnitalSubsemiringClass.toNonUnitalSemiring (closure s) with + mul_comm := fun ⟨_, h₁⟩ ⟨_, h₂⟩ ↦ + have := closure_le_centralizer_centralizer s + Subtype.ext <| Set.centralizer_centralizer_comm_of_comm hcomm _ (this h₁) _ (this h₂) } + variable [NonUnitalNonAssocSemiring S] theorem mem_map_equiv {f : R ≃+* S} {K : NonUnitalSubsemiring R} {x : S} : From 8925b43f278bde9f3f23715e1c7d1ec3353b6215 Mon Sep 17 00:00:00 2001 From: Yongle Hu Date: Sun, 17 Nov 2024 15:35:08 +0000 Subject: [PATCH 028/829] chore(NumberTheory/RamificationInertia): move file RamificationInertia into folder RamificationInertia (#19140) Move the file `RamificationInertia.lean` into folder `NumberTheory/RamificationInertia` since we upstream the ramification theory in Galois extensions in #19141. Moves: - Mathlib.NumberTheory.RamificationInertia -> Mathlib.NumberTheory.RamificationInertia.Basic - [x] depends on: #19138 Co-authored-by: Hu Yongle Co-authored-by: Yongle Hu <140475041+mbkybky@users.noreply.github.com> --- Mathlib.lean | 2 +- .../Basic.lean} | 0 2 files changed, 1 insertion(+), 1 deletion(-) rename Mathlib/NumberTheory/{RamificationInertia.lean => RamificationInertia/Basic.lean} (100%) diff --git a/Mathlib.lean b/Mathlib.lean index 4902e66c634db..8d465054d65ab 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3823,7 +3823,7 @@ import Mathlib.NumberTheory.PrimeCounting import Mathlib.NumberTheory.PrimesCongruentOne import Mathlib.NumberTheory.Primorial import Mathlib.NumberTheory.PythagoreanTriples -import Mathlib.NumberTheory.RamificationInertia +import Mathlib.NumberTheory.RamificationInertia.Basic import Mathlib.NumberTheory.Rayleigh import Mathlib.NumberTheory.SiegelsLemma import Mathlib.NumberTheory.SmoothNumbers diff --git a/Mathlib/NumberTheory/RamificationInertia.lean b/Mathlib/NumberTheory/RamificationInertia/Basic.lean similarity index 100% rename from Mathlib/NumberTheory/RamificationInertia.lean rename to Mathlib/NumberTheory/RamificationInertia/Basic.lean From 14c7a8e80b854559a1a4aa2144d3f510feaa7ddf Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Sun, 17 Nov 2024 16:01:25 +0000 Subject: [PATCH 029/829] feat(CategoryTheory/Enriched): functor categories are enriched (#18009) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Let `C` be a category that is enriched over a monoidal category `V` in such a way that the category structure and the enriched category structure are compatible. Then, if `J` is a category and that `V` has certain limits, then the functor category `J ⥤ C` is also enriched over `V`. (Plan: using #17326, we may use this for `C := C` closed monoidal in order to show that a category of functors `J ⥤ C` to a monoidal category is enriched over `C`, and, by applying this to all `Under X` categories for `X : C`, it should follow that `J ⥤ C` is also closed monoidal. This should give a more explicit approach as compared to #16067.) Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- Mathlib.lean | 1 + .../Enriched/FunctorCategory.lean | 219 ++++++++++++++++++ Mathlib/CategoryTheory/Enriched/Ordinary.lean | 6 +- 3 files changed, 223 insertions(+), 3 deletions(-) create mode 100644 Mathlib/CategoryTheory/Enriched/FunctorCategory.lean diff --git a/Mathlib.lean b/Mathlib.lean index 8d465054d65ab..04e0f71870e2a 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1650,6 +1650,7 @@ import Mathlib.CategoryTheory.Elementwise import Mathlib.CategoryTheory.Endofunctor.Algebra import Mathlib.CategoryTheory.Endomorphism import Mathlib.CategoryTheory.Enriched.Basic +import Mathlib.CategoryTheory.Enriched.FunctorCategory import Mathlib.CategoryTheory.Enriched.Ordinary import Mathlib.CategoryTheory.EpiMono import Mathlib.CategoryTheory.EqToHom diff --git a/Mathlib/CategoryTheory/Enriched/FunctorCategory.lean b/Mathlib/CategoryTheory/Enriched/FunctorCategory.lean new file mode 100644 index 0000000000000..5d53ac120e5ee --- /dev/null +++ b/Mathlib/CategoryTheory/Enriched/FunctorCategory.lean @@ -0,0 +1,219 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.CategoryTheory.Enriched.Ordinary +import Mathlib.CategoryTheory.Functor.Category +import Mathlib.CategoryTheory.Limits.Shapes.End + +/-! +# Functor categories are enriched + +If `C` is a `V`-enriched ordinary category, then `J ⥤ C` is also +a `V`-enriched ordinary category, provided `C` has suitable limits. + +-/ + +universe v₁ v₂ v₃ u₁ u₂ u₃ + +namespace CategoryTheory.Enriched.FunctorCategory + +open Category MonoidalCategory Limits + +variable (V : Type u₁) [Category.{v₁} V] [MonoidalCategory V] + {C : Type u₂} [Category.{v₂} C] + {J : Type u₃} [Category.{v₃} J] [EnrichedOrdinaryCategory V C] + +variable (F₁ F₂ F₃ F₄ : J ⥤ C) + +/-- Given two functors `F₁` and `F₂` from a category `J` to a `V`-enriched +ordinary category `C`, this is the diagram `Jᵒᵖ ⥤ J ⥤ V` whose end shall be +the `V`-morphisms in `J ⥤ V` from `F₁` to `F₂`. -/ +@[simps!] +def diagram : Jᵒᵖ ⥤ J ⥤ V := F₁.op ⋙ eHomFunctor V C ⋙ (whiskeringLeft J C V).obj F₂ + +/-- The condition that the end `diagram V F₁ F₂` exists, see `enrichedHom`. -/ +abbrev HasEnrichedHom := HasEnd (diagram V F₁ F₂) + +section + +variable [HasEnrichedHom V F₁ F₂] + +/-- The `V`-enriched hom from `F₁` to `F₂` when `F₁` and `F₂` are functors `J ⥤ C` +and `C` is a `V`-enriched category. -/ +noncomputable abbrev enrichedHom : V := end_ (diagram V F₁ F₂) + +/-- The projection `enrichedHom V F₁ F₂ ⟶ F₁.obj j ⟶[V] F₂.obj j` in the category `V` +for any `j : J` when `F₁` and `F₂` are functors `J ⥤ C` and `C` is a `V`-enriched category. -/ +noncomputable abbrev enrichedHomπ (j : J) : enrichedHom V F₁ F₂ ⟶ F₁.obj j ⟶[V] F₂.obj j := + end_.π _ j + +@[reassoc] +lemma enrichedHom_condition {i j : J} (f : i ⟶ j) : + enrichedHomπ V F₁ F₂ i ≫ (ρ_ _).inv ≫ + _ ◁ (eHomEquiv V) (F₂.map f) ≫ eComp V _ _ _ = + enrichedHomπ V F₁ F₂ j ≫ (λ_ _).inv ≫ + (eHomEquiv V) (F₁.map f) ▷ _ ≫ eComp V _ _ _ := + end_.condition (diagram V F₁ F₂) f + +variable {F₁ F₂} + +/-- Given functors `F₁` and `F₂` in `J ⥤ C`, where `C` is a `V`-enriched ordinary category, +this is the isomorphism `(F₁ ⟶ F₂) ≃ (𝟙_ V ⟶ enrichedHom V F₁ F₂)` in the category `V`. -/ +noncomputable def homEquiv : (F₁ ⟶ F₂) ≃ (𝟙_ V ⟶ enrichedHom V F₁ F₂) where + toFun τ := end_.lift (fun j ↦ eHomEquiv V (τ.app j)) (fun i j f ↦ by + trans eHomEquiv V (τ.app i ≫ F₂.map f) + · dsimp + simp only [eHomEquiv_comp, tensorHom_def_assoc, MonoidalCategory.whiskerRight_id, + ← unitors_equal, assoc, Iso.inv_hom_id_assoc, eHomWhiskerLeft] + · dsimp + simp only [← NatTrans.naturality, eHomEquiv_comp, tensorHom_def', id_whiskerLeft, + assoc, Iso.inv_hom_id_assoc, eHomWhiskerRight]) + invFun g := + { app := fun j ↦ (eHomEquiv V).symm (g ≫ end_.π _ j) + naturality := fun i j f ↦ (eHomEquiv V).injective (by + dsimp + simp only [eHomEquiv_comp, Equiv.apply_symm_apply, Iso.cancel_iso_inv_left] + conv_rhs => + rw [tensorHom_def_assoc, MonoidalCategory.whiskerRight_id_assoc, assoc, + enrichedHom_condition V F₁ F₂ f] + conv_lhs => + rw [tensorHom_def'_assoc, MonoidalCategory.whiskerLeft_comp_assoc, + id_whiskerLeft_assoc, id_whiskerLeft_assoc, Iso.inv_hom_id_assoc, unitors_equal]) } + left_inv τ := by aesop + right_inv g := by aesop + +@[reassoc (attr := simp)] +lemma homEquiv_apply_π (τ : F₁ ⟶ F₂) (j : J) : + homEquiv V τ ≫ enrichedHomπ V _ _ j = eHomEquiv V (τ.app j) := by + simp [homEquiv] + +end + +section + +variable [HasEnrichedHom V F₁ F₁] + +/-- The identity for the `V`-enrichment of the category `J ⥤ C` over `V`. -/ +noncomputable def enrichedId : 𝟙_ V ⟶ enrichedHom V F₁ F₁ := homEquiv _ (𝟙 F₁) + +@[reassoc (attr := simp)] +lemma enrichedId_π (j : J) : enrichedId V F₁ ≫ end_.π _ j = eId V (F₁.obj j) := by + simp [enrichedId] + +@[simp] +lemma homEquiv_id : homEquiv V (𝟙 F₁) = enrichedId V F₁ := rfl + +end + +section + +variable [HasEnrichedHom V F₁ F₂] [HasEnrichedHom V F₂ F₃] [HasEnrichedHom V F₁ F₃] + +/-- The composition for the `V`-enrichment of the category `J ⥤ C` over `V`. -/ +noncomputable def enrichedComp : enrichedHom V F₁ F₂ ⊗ enrichedHom V F₂ F₃ ⟶ enrichedHom V F₁ F₃ := + end_.lift (fun j ↦ (end_.π _ j ⊗ end_.π _ j) ≫ eComp V _ _ _) (fun i j f ↦ by + dsimp + trans (end_.π (diagram V F₁ F₂) i ⊗ end_.π (diagram V F₂ F₃) j) ≫ + (ρ_ _).inv ▷ _ ≫ (_ ◁ (eHomEquiv V (F₂.map f))) ▷ _ ≫ eComp V _ (F₂.obj i) _ ▷ _ ≫ + eComp V _ (F₂.obj j) _ + · have := end_.condition (diagram V F₂ F₃) f + dsimp [eHomWhiskerLeft, eHomWhiskerRight] at this ⊢ + conv_lhs => rw [assoc, tensorHom_def_assoc] + conv_rhs => + rw [tensorHom_def_assoc, whisker_assoc_assoc, e_assoc, + triangle_assoc_comp_right_inv_assoc, ← MonoidalCategory.whiskerLeft_comp_assoc, + ← MonoidalCategory.whiskerLeft_comp_assoc, ← MonoidalCategory.whiskerLeft_comp_assoc, + assoc, assoc, ← this, MonoidalCategory.whiskerLeft_comp_assoc, + MonoidalCategory.whiskerLeft_comp_assoc, MonoidalCategory.whiskerLeft_comp_assoc, + ← e_assoc, whiskerLeft_rightUnitor_inv_assoc, associator_inv_naturality_right_assoc, + Iso.hom_inv_id_assoc, whisker_exchange_assoc, MonoidalCategory.whiskerRight_id_assoc, + Iso.inv_hom_id_assoc] + · have := end_.condition (diagram V F₁ F₂) f + dsimp [eHomWhiskerLeft, eHomWhiskerRight] at this ⊢ + conv_lhs => + rw [tensorHom_def'_assoc, ← comp_whiskerRight_assoc, + ← comp_whiskerRight_assoc, ← comp_whiskerRight_assoc, + assoc, assoc, this, comp_whiskerRight_assoc, comp_whiskerRight_assoc, + comp_whiskerRight_assoc, leftUnitor_inv_whiskerRight_assoc, + ← associator_inv_naturality_left_assoc, ← e_assoc', + Iso.inv_hom_id_assoc, ← whisker_exchange_assoc, id_whiskerLeft_assoc, + Iso.inv_hom_id_assoc] + conv_rhs => rw [assoc, tensorHom_def'_assoc]) + +@[reassoc (attr := simp)] +lemma enrichedComp_π (j : J) : + enrichedComp V F₁ F₂ F₃ ≫ end_.π _ j = + (end_.π (diagram V F₁ F₂) j ⊗ end_.π (diagram V F₂ F₃) j) ≫ eComp V _ _ _ := by + simp [enrichedComp] + +variable {F₁ F₂ F₃} + +@[reassoc] +lemma homEquiv_comp (f : F₁ ⟶ F₂) (g : F₂ ⟶ F₃) : + (homEquiv V) (f ≫ g) = (λ_ (𝟙_ V)).inv ≫ ((homEquiv V) f ⊗ (homEquiv V) g) ≫ + enrichedComp V F₁ F₂ F₃ := by + ext j + simp only [homEquiv_apply_π, NatTrans.comp_app, eHomEquiv_comp, assoc, + enrichedComp_π, Functor.op_obj, ← tensor_comp_assoc] + +end + +@[reassoc (attr := simp)] +lemma enriched_id_comp [HasEnrichedHom V F₁ F₁] [HasEnrichedHom V F₁ F₂] : + (λ_ (enrichedHom V F₁ F₂)).inv ≫ enrichedId V F₁ ▷ enrichedHom V F₁ F₂ ≫ + enrichedComp V F₁ F₁ F₂ = 𝟙 _ := by + ext j + rw [assoc, assoc, enrichedComp_π, id_comp, tensorHom_def, assoc, + ← comp_whiskerRight_assoc, enrichedId_π, ← whisker_exchange_assoc, + id_whiskerLeft, assoc, assoc, Iso.inv_hom_id_assoc] + dsimp + rw [e_id_comp, comp_id] + +@[reassoc (attr := simp)] +lemma enriched_comp_id [HasEnrichedHom V F₁ F₂] [HasEnrichedHom V F₂ F₂] : + (ρ_ (enrichedHom V F₁ F₂)).inv ≫ enrichedHom V F₁ F₂ ◁ enrichedId V F₂ ≫ + enrichedComp V F₁ F₂ F₂ = 𝟙 _ := by + ext j + rw [assoc, assoc, enrichedComp_π, id_comp, tensorHom_def', assoc, + ← MonoidalCategory.whiskerLeft_comp_assoc, enrichedId_π, + whisker_exchange_assoc, MonoidalCategory.whiskerRight_id, assoc, assoc, + Iso.inv_hom_id_assoc] + dsimp + rw [e_comp_id, comp_id] + +@[reassoc] +lemma enriched_assoc [HasEnrichedHom V F₁ F₂] [HasEnrichedHom V F₁ F₃] [HasEnrichedHom V F₁ F₄] + [HasEnrichedHom V F₂ F₃] [HasEnrichedHom V F₂ F₄] [HasEnrichedHom V F₃ F₄] : + (α_ (enrichedHom V F₁ F₂) (enrichedHom V F₂ F₃) (enrichedHom V F₃ F₄)).inv ≫ + enrichedComp V F₁ F₂ F₃ ▷ enrichedHom V F₃ F₄ ≫ enrichedComp V F₁ F₃ F₄ = + enrichedHom V F₁ F₂ ◁ enrichedComp V F₂ F₃ F₄ ≫ enrichedComp V F₁ F₂ F₄ := by + ext j + conv_lhs => + rw [assoc, assoc, enrichedComp_π, + tensorHom_def_assoc, ← comp_whiskerRight_assoc, enrichedComp_π, + comp_whiskerRight_assoc, ← whisker_exchange_assoc, + ← whisker_exchange_assoc, ← tensorHom_def'_assoc, ← associator_inv_naturality_assoc] + conv_rhs => + rw [assoc, enrichedComp_π, tensorHom_def'_assoc, ← MonoidalCategory.whiskerLeft_comp_assoc, + enrichedComp_π, MonoidalCategory.whiskerLeft_comp_assoc, whisker_exchange_assoc, + whisker_exchange_assoc, ← tensorHom_def_assoc] + dsimp + rw [e_assoc] + +variable (J C) + +/-- If `C` is a `V`-enriched ordinary category, and `C` has suitable limits, +then `J ⥤ C` is also a `V`-enriched ordinary category. -/ +noncomputable def enrichedOrdinaryCategory [∀ (F₁ F₂ : J ⥤ C), HasEnrichedHom V F₁ F₂] : + EnrichedOrdinaryCategory V (J ⥤ C) where + Hom F₁ F₂ := enrichedHom V F₁ F₂ + id F := enrichedId V F + comp F₁ F₂ F₃ := enrichedComp V F₁ F₂ F₃ + assoc _ _ _ _ := enriched_assoc _ _ _ _ _ + homEquiv := homEquiv V + homEquiv_id _ := homEquiv_id V _ + homEquiv_comp f g := homEquiv_comp V f g + +end CategoryTheory.Enriched.FunctorCategory diff --git a/Mathlib/CategoryTheory/Enriched/Ordinary.lean b/Mathlib/CategoryTheory/Enriched/Ordinary.lean index e7081037f3137..85c4db8068474 100644 --- a/Mathlib/CategoryTheory/Enriched/Ordinary.lean +++ b/Mathlib/CategoryTheory/Enriched/Ordinary.lean @@ -62,7 +62,7 @@ lemma eHomEquiv_comp {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) : EnrichedOrdinaryCategory.homEquiv_comp _ _ /-- The morphism `(X' ⟶[V] Y) ⟶ (X ⟶[V] Y)` induced by a morphism `X ⟶ X'`. -/ -noncomputable def eHomWhiskerRight {X X' : C} (f : X ⟶ X') (Y : C) : +def eHomWhiskerRight {X X' : C} (f : X ⟶ X') (Y : C) : (X' ⟶[V] Y) ⟶ (X ⟶[V] Y) := (λ_ _).inv ≫ eHomEquiv V f ▷ _ ≫ eComp V X X' Y @@ -82,7 +82,7 @@ lemma eHomWhiskerRight_comp {X X' X'' : C} (f : X ⟶ X') (f' : X' ⟶ X'') (Y : ← whisker_exchange_assoc, id_whiskerLeft_assoc, Iso.inv_hom_id_assoc] /-- The morphism `(X ⟶[V] Y) ⟶ (X ⟶[V] Y')` induced by a morphism `Y ⟶ Y'`. -/ -noncomputable def eHomWhiskerLeft (X : C) {Y Y' : C} (g : Y ⟶ Y') : +def eHomWhiskerLeft (X : C) {Y Y' : C} (g : Y ⟶ Y') : (X ⟶[V] Y) ⟶ (X ⟶[V] Y') := (ρ_ _).inv ≫ _ ◁ eHomEquiv V g ≫ eComp V X Y Y' @@ -121,7 +121,7 @@ attribute [local simp] eHom_whisker_exchange variable (C) in /-- The bifunctor `Cᵒᵖ ⥤ C ⥤ V` which sends `X : Cᵒᵖ` and `Y : C` to `X ⟶[V] Y`. -/ @[simps] -noncomputable def eHomFunctor : Cᵒᵖ ⥤ C ⥤ V where +def eHomFunctor : Cᵒᵖ ⥤ C ⥤ V where obj X := { obj := fun Y => X.unop ⟶[V] Y map := fun φ => eHomWhiskerLeft V X.unop φ } From 62ee20ffeeb616f2b1e9cfd578d1d8252210275c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 17 Nov 2024 16:01:26 +0000 Subject: [PATCH 030/829] feat(Polynomial): composing with the constant polynomial (#19090) ... and a point-free version of `coeff_map` From GrowthInGroups --- Mathlib/Algebra/Polynomial/Eval/Coeff.lean | 3 +++ Mathlib/Algebra/Polynomial/Eval/Defs.lean | 3 +++ 2 files changed, 6 insertions(+) diff --git a/Mathlib/Algebra/Polynomial/Eval/Coeff.lean b/Mathlib/Algebra/Polynomial/Eval/Coeff.lean index 882cb4347f8ff..cd0d1add2ccd8 100644 --- a/Mathlib/Algebra/Polynomial/Eval/Coeff.lean +++ b/Mathlib/Algebra/Polynomial/Eval/Coeff.lean @@ -80,6 +80,9 @@ theorem coeff_map (n : ℕ) : coeff (p.map f) n = f (coeff p n) := by simp only [RingHom.coe_comp, Function.comp, coeff_C_mul_X_pow] split_ifs <;> simp [f.map_zero] +lemma coeff_map_eq_comp (p : R[X]) (f : R →+* S) : (p.map f).coeff = f ∘ p.coeff := by + ext n; exact coeff_map .. + theorem map_map [Semiring T] (g : S →+* T) (p : R[X]) : (p.map f).map g = p.map (g.comp f) := ext (by simp [coeff_map]) diff --git a/Mathlib/Algebra/Polynomial/Eval/Defs.lean b/Mathlib/Algebra/Polynomial/Eval/Defs.lean index a7850951cd053..af51045eb7816 100644 --- a/Mathlib/Algebra/Polynomial/Eval/Defs.lean +++ b/Mathlib/Algebra/Polynomial/Eval/Defs.lean @@ -560,6 +560,9 @@ protected theorem map_ofNat (n : ℕ) [n.AtLeastTwo] : theorem map_dvd (f : R →+* S) {x y : R[X]} : x ∣ y → x.map f ∣ y.map f := (mapRingHom f).map_dvd +lemma mapRingHom_comp_C {R S} [CommRing R] [CommRing S] (f : R →+* S) : + (mapRingHom f).comp C = C.comp f := by ext; simp + theorem eval₂_eq_eval_map {x : S} : p.eval₂ f x = (p.map f).eval x := by induction p using Polynomial.induction_on' with | h_add p q hp hq => From acffbb9f7facd8b48378c4ba8e4a717614ae55ee Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 17 Nov 2024 16:39:35 +0000 Subject: [PATCH 031/829] chore: deduplicate `BddAbove.mul` (#19161) --- Mathlib.lean | 1 - .../Algebra/Order/Group/Pointwise/Bounds.lean | 2 ++ Mathlib/Data/Set/Pointwise/BoundedMul.lean | 28 ------------------- 3 files changed, 2 insertions(+), 29 deletions(-) delete mode 100644 Mathlib/Data/Set/Pointwise/BoundedMul.lean diff --git a/Mathlib.lean b/Mathlib.lean index 04e0f71870e2a..928ce53badc19 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2800,7 +2800,6 @@ import Mathlib.Data.Set.Opposite import Mathlib.Data.Set.Pairwise.Basic import Mathlib.Data.Set.Pairwise.Lattice import Mathlib.Data.Set.Pointwise.BigOperators -import Mathlib.Data.Set.Pointwise.BoundedMul import Mathlib.Data.Set.Pointwise.Finite import Mathlib.Data.Set.Pointwise.Interval import Mathlib.Data.Set.Pointwise.Iterate diff --git a/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean b/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean index d7e523a2a6595..ceb329876b8e6 100644 --- a/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean +++ b/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean @@ -48,6 +48,8 @@ lemma BddAbove.mul (hs : BddAbove s) (ht : BddAbove t) : BddAbove (s * t) := lemma BddBelow.mul (hs : BddBelow s) (ht : BddBelow t) : BddBelow (s * t) := (Nonempty.mul hs ht).mono (subset_lowerBounds_mul s t) +@[to_additive (attr := deprecated (since := "2024-11-13"))] alias Set.BddAbove.mul := BddAbove.mul + @[to_additive] lemma BddAbove.range_mul (hf : BddAbove (range f)) (hg : BddAbove (range g)) : BddAbove (range fun i ↦ f i * g i) := diff --git a/Mathlib/Data/Set/Pointwise/BoundedMul.lean b/Mathlib/Data/Set/Pointwise/BoundedMul.lean deleted file mode 100644 index 2d630b685665e..0000000000000 --- a/Mathlib/Data/Set/Pointwise/BoundedMul.lean +++ /dev/null @@ -1,28 +0,0 @@ -/- -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.Basic -import Mathlib.Algebra.Order.Monoid.Defs - -/-! - -# Pointwise multiplication of sets preserves boundedness above. - -TODO: Can be combined with future results about pointwise multiplication on sets that use -ordered algebraic classes. - --/ - -variable {α : Type*} - -namespace Set - -open Pointwise - -@[to_additive] -theorem BddAbove.mul [OrderedCommMonoid α] {A B : Set α} (hA : BddAbove A) (hB : BddAbove B) : - BddAbove (A * B) := - hA.image2 (fun _ _ _ h ↦ mul_le_mul_right' h _) (fun _ _ _ h ↦ mul_le_mul_left' h _) hB -end Set From 561dec6daae81e2640ff33e99162cef451178412 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 17 Nov 2024 16:59:58 +0000 Subject: [PATCH 032/829] =?UTF-8?q?feat:=20`(WithZero=20=CE=B1)=CB=A3`=20i?= =?UTF-8?q?s=20isomorphic=20to=20`=CE=B1`=20as=20an=20ordered=20group=20(#?= =?UTF-8?q?18986)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From FLT --- Mathlib/Algebra/Order/GroupWithZero/Canonical.lean | 8 ++++++++ Mathlib/Algebra/Order/Hom/Monoid.lean | 6 +++++- 2 files changed, 13 insertions(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean index 8140a3b819843..e0e51858c4b0b 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean @@ -299,6 +299,14 @@ lemma zero_eq_bot : (0 : WithZero α) = ⊥ := rfl theorem coe_le_iff {x : WithZero α} : (a : WithZero α) ≤ x ↔ ∃ b : α, x = b ∧ a ≤ b := WithBot.coe_le_iff +@[simp] lemma unzero_le_unzero {a b : WithZero α} (ha hb) : + unzero (x := a) ha ≤ unzero (x := b) hb ↔ a ≤ b := by + -- TODO: Fix `lift` so that it doesn't try to clear the hypotheses I give it when it is + -- impossible to do so. See https://github.com/leanprover-community/mathlib4/issues/19160 + lift a to α using id ha + lift b to α using id hb + simp + instance mulLeftMono [Mul α] [MulLeftMono α] : MulLeftMono (WithZero α) := by refine ⟨fun a b c hbc => ?_⟩ diff --git a/Mathlib/Algebra/Order/Hom/Monoid.lean b/Mathlib/Algebra/Order/Hom/Monoid.lean index e2cca36e6ab51..5fa19423317a2 100644 --- a/Mathlib/Algebra/Order/Hom/Monoid.lean +++ b/Mathlib/Algebra/Order/Hom/Monoid.lean @@ -841,4 +841,8 @@ end LinearOrderedCommMonoidWithZero end OrderMonoidWithZeroHom -/- See module docstring for details. -/ +/-- Any ordered group is isomorphic to the units of itself adjoined with `0`. -/ +@[simps! toFun] +def OrderMonoidIso.unitsWithZero {α : Type*} [Group α] [Preorder α] : (WithZero α)ˣ ≃*o α where + toMulEquiv := WithZero.unitsWithZeroEquiv + map_le_map_iff' {a b} := by simp [WithZero.unitsWithZeroEquiv] From 67997fa092e6fa0086ff6d6e14f6e0ed9b688db3 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 17 Nov 2024 16:59:59 +0000 Subject: [PATCH 033/829] refactor(ENat.Lattice): ENat.sSup_mem_of_nonempty_of_lt_top naming convention (#19147) The `of_nonempty` should not be capitalized. --- Mathlib/Data/ENat/Lattice.lean | 4 ++-- Mathlib/Order/KrullDimension.lean | 2 +- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/Data/ENat/Lattice.lean b/Mathlib/Data/ENat/Lattice.lean index 376dff287b55f..d56040d7d9e21 100644 --- a/Mathlib/Data/ENat/Lattice.lean +++ b/Mathlib/Data/ENat/Lattice.lean @@ -105,12 +105,12 @@ lemma finite_of_sSup_lt_top (h : sSup s < ⊤) : s.Finite := by simp only [top_le_iff] exact sSup_eq_top_of_infinite h -lemma sSup_mem_of_Nonempty_of_lt_top [Nonempty s] (hs' : sSup s < ⊤) : sSup s ∈ s := +lemma sSup_mem_of_nonempty_of_lt_top [Nonempty s] (hs' : sSup s < ⊤) : sSup s ∈ s := Nonempty.csSup_mem nonempty_of_nonempty_subtype (finite_of_sSup_lt_top hs') lemma exists_eq_iSup_of_lt_top [Nonempty ι] (h : ⨆ i, f i < ⊤) : ∃ i, f i = ⨆ i, f i := - sSup_mem_of_Nonempty_of_lt_top h + sSup_mem_of_nonempty_of_lt_top h lemma exists_eq_iSup₂_of_lt_top {ι₁ ι₂ : Type*} {f : ι₁ → ι₂ → ℕ∞} [Nonempty ι₁] [Nonempty ι₂] (h : ⨆ i, ⨆ j, f i j < ⊤) : ∃ i j, f i j = ⨆ i, ⨆ j, f i j := by diff --git a/Mathlib/Order/KrullDimension.lean b/Mathlib/Order/KrullDimension.lean index bd5570bd4816a..7750f256b126e 100644 --- a/Mathlib/Order/KrullDimension.lean +++ b/Mathlib/Order/KrullDimension.lean @@ -194,7 +194,7 @@ lemma height_orderIso (f : α ≃o β) (x : α) : height (f x) = height x := by private lemma exists_eq_iSup_of_iSup_eq_coe {α : Type*} [Nonempty α] {f : α → ℕ∞} {n : ℕ} (h : (⨆ x, f x) = n) : ∃ x, f x = n := by - obtain ⟨x, hx⟩ := ENat.sSup_mem_of_Nonempty_of_lt_top (h ▸ ENat.coe_lt_top _) + obtain ⟨x, hx⟩ := ENat.sSup_mem_of_nonempty_of_lt_top (h ▸ ENat.coe_lt_top _) use x simpa [hx] using h From 34691d71065ed8782d32ef4bd34bbb2f0d41a985 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Sun, 17 Nov 2024 17:00:01 +0000 Subject: [PATCH 034/829] chore: use `RingHom.ker` for `Polynomial.ker_mapRingHom` (#19157) --- Mathlib/RingTheory/Polynomial/Basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/RingTheory/Polynomial/Basic.lean b/Mathlib/RingTheory/Polynomial/Basic.lean index 306da268333af..99c10175f0d49 100644 --- a/Mathlib/RingTheory/Polynomial/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Basic.lean @@ -577,9 +577,9 @@ theorem mem_map_C_iff {I : Ideal R} {f : R[X]} : exact (I.map C : Ideal R[X]).mul_mem_left _ (mem_map_of_mem _ (hf n)) theorem _root_.Polynomial.ker_mapRingHom (f : R →+* S) : - LinearMap.ker (Polynomial.mapRingHom f).toSemilinearMap = f.ker.map (C : R →+* R[X]) := by + RingHom.ker (Polynomial.mapRingHom f) = f.ker.map (C : R →+* R[X]) := by ext - simp only [LinearMap.mem_ker, RingHom.toSemilinearMap_apply, coe_mapRingHom] + simp only [RingHom.mem_ker, coe_mapRingHom] rw [mem_map_C_iff, Polynomial.ext_iff] simp [RingHom.mem_ker] From 8161e351889398492eaaf96db0cad75c0b048fa8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 17 Nov 2024 17:00:02 +0000 Subject: [PATCH 035/829] chore(Tactic/Lift): don't use `True`/`False` where a `Bool` is expected (#19159) --- Mathlib/Tactic/Lift.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Tactic/Lift.lean b/Mathlib/Tactic/Lift.lean index c547533cbf31a..e24fa486a5c55 100644 --- a/Mathlib/Tactic/Lift.lean +++ b/Mathlib/Tactic/Lift.lean @@ -162,20 +162,20 @@ def Lift.main (e t : TSyntax `term) (hUsing : Option (TSyntax `term)) if hUsing.isNone then withMainContext <| setGoals (prf.mvarId! :: (← getGoals)) elab_rules : tactic - | `(tactic| lift $e to $t $[using $h]?) => withMainContext <| Lift.main e t h none none False + | `(tactic| lift $e to $t $[using $h]?) => withMainContext <| Lift.main e t h none none false elab_rules : tactic | `(tactic| lift $e to $t $[using $h]? - with $newVarName) => withMainContext <| Lift.main e t h newVarName none False + with $newVarName) => withMainContext <| Lift.main e t h newVarName none false elab_rules : tactic | `(tactic| lift $e to $t $[using $h]? - with $newVarName $newEqName) => withMainContext <| Lift.main e t h newVarName newEqName False + with $newVarName $newEqName) => withMainContext <| Lift.main e t h newVarName newEqName false elab_rules : tactic | `(tactic| lift $e to $t $[using $h]? with $newVarName $newEqName $newPrfName) => withMainContext do - if h.isNone then Lift.main e t h newVarName newEqName False + if h.isNone then Lift.main e t h newVarName newEqName false else let some h := h | unreachable! - if h.raw == newPrfName then Lift.main e t h newVarName newEqName True - else Lift.main e t h newVarName newEqName False + if h.raw == newPrfName then Lift.main e t h newVarName newEqName true + else Lift.main e t h newVarName newEqName false end Mathlib.Tactic From 917c0790c6a67febe3ccaa35a68da9c189c1c757 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Sun, 17 Nov 2024 17:51:15 +0000 Subject: [PATCH 036/829] refactor(Algebra/Polynomial/Lifts): Golf some proofs (#18796) This PR golfs some proofs in `Algebra/Polynomial/Lifts.lean`. Co-authored-by: Thomas Browning --- Mathlib/Algebra/Polynomial/Lifts.lean | 96 +++++++-------------------- 1 file changed, 25 insertions(+), 71 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/Lifts.lean b/Mathlib/Algebra/Polynomial/Lifts.lean index 0ded04695c7fd..622ecb58735da 100644 --- a/Mathlib/Algebra/Polynomial/Lifts.lean +++ b/Mathlib/Algebra/Polynomial/Lifts.lean @@ -123,63 +123,26 @@ section LiftDeg theorem monomial_mem_lifts_and_degree_eq {s : S} {n : ℕ} (hl : monomial n s ∈ lifts f) : ∃ q : R[X], map f q = monomial n s ∧ q.degree = (monomial n s).degree := by - by_cases hzero : s = 0 - · use 0 - simp only [hzero, degree_zero, eq_self_iff_true, and_self_iff, monomial_zero_right, - Polynomial.map_zero] - rw [lifts_iff_set_range] at hl - obtain ⟨q, hq⟩ := hl - replace hq := (ext_iff.1 hq) n - have hcoeff : f (q.coeff n) = s := by - rwa [coeff_map, coeff_monomial_same] at hq - use monomial n (q.coeff n) - constructor - · simp only [hcoeff, map_monomial] - have hqzero : q.coeff n ≠ 0 := by - intro habs - simp only [habs, RingHom.map_zero] at hcoeff - exact hzero hcoeff.symm - rw [← C_mul_X_pow_eq_monomial, ← C_mul_X_pow_eq_monomial] - simp only [hzero, hqzero, Ne, not_false_iff, degree_C_mul_X_pow] + rcases eq_or_ne s 0 with rfl | h + · exact ⟨0, by simp⟩ + obtain ⟨a, rfl⟩ := coeff_monomial_same n s ▸ (monomial n s).lifts_iff_coeff_lifts.mp hl n + refine ⟨monomial n a, map_monomial f, ?_⟩ + rw [degree_monomial, degree_monomial n h] + exact mt (fun ha ↦ ha ▸ map_zero f) h /-- A polynomial lifts if and only if it can be lifted to a polynomial of the same degree. -/ theorem mem_lifts_and_degree_eq {p : S[X]} (hlifts : p ∈ lifts f) : ∃ q : R[X], map f q = p ∧ q.degree = p.degree := by - generalize hd : p.natDegree = d - revert hd p - induction' d using Nat.strong_induction_on with n hn - intros p hlifts hdeg - by_cases erase_zero : p.eraseLead = 0 - · rw [← eraseLead_add_monomial_natDegree_leadingCoeff p, erase_zero, zero_add, leadingCoeff] - exact - monomial_mem_lifts_and_degree_eq - (monomial_mem_lifts p.natDegree ((lifts_iff_coeff_lifts p).1 hlifts p.natDegree)) - have deg_erase := Or.resolve_right (eraseLead_natDegree_lt_or_eraseLead_eq_zero p) erase_zero - have pzero : p ≠ 0 := by - intro habs - exfalso - rw [habs, eraseLead_zero, eq_self_iff_true, not_true] at erase_zero - exact erase_zero - have lead_zero : p.coeff p.natDegree ≠ 0 := by - rw [← leadingCoeff, Ne, leadingCoeff_eq_zero]; exact pzero - obtain ⟨lead, hlead⟩ := - monomial_mem_lifts_and_degree_eq - (monomial_mem_lifts p.natDegree ((lifts_iff_coeff_lifts p).1 hlifts p.natDegree)) - have deg_lead : lead.degree = p.natDegree := by - rw [hlead.2, ← C_mul_X_pow_eq_monomial, degree_C_mul_X_pow p.natDegree lead_zero] - rw [hdeg] at deg_erase - obtain ⟨erase, herase⟩ := - hn p.eraseLead.natDegree deg_erase (erase_mem_lifts p.natDegree hlifts) - (refl p.eraseLead.natDegree) - use erase + lead - constructor - · simp only [hlead, herase, Polynomial.map_add] - rw [← eraseLead, ← leadingCoeff] - rw [eraseLead_add_monomial_natDegree_leadingCoeff p] - rw [degree_eq_natDegree pzero, ← deg_lead] - apply degree_add_eq_right_of_degree_lt - rw [herase.2, deg_lead, ← degree_eq_natDegree pzero] - exact degree_erase_lt pzero + rw [lifts_iff_coeff_lifts] at hlifts + let g : ℕ → R := fun k ↦ (hlifts k).choose + have hg : ∀ k, f (g k) = p.coeff k := fun k ↦ (hlifts k).choose_spec + let q : R[X] := ∑ k ∈ p.support, monomial k (g k) + have hq : map f q = p := by simp_rw [q, Polynomial.map_sum, map_monomial, hg, ← as_sum_support] + have hq' : q.support = p.support := by + simp_rw [Finset.ext_iff, mem_support_iff, q, finset_sum_coeff, coeff_monomial, + Finset.sum_ite_eq', ite_ne_right_iff, mem_support_iff, and_iff_left_iff_imp, not_imp_not] + exact fun k h ↦ by rw [← hg, h, map_zero] + exact ⟨q, hq, congrArg Finset.max hq'⟩ end LiftDeg @@ -189,24 +152,15 @@ section Monic of the same degree. -/ theorem lifts_and_degree_eq_and_monic [Nontrivial S] {p : S[X]} (hlifts : p ∈ lifts f) (hp : p.Monic) : ∃ q : R[X], map f q = p ∧ q.degree = p.degree ∧ q.Monic := by - cases' subsingleton_or_nontrivial R with hR hR - · obtain ⟨q, hq⟩ := mem_lifts_and_degree_eq hlifts - exact ⟨q, hq.1, hq.2, monic_of_subsingleton _⟩ - have H : erase p.natDegree p + X ^ p.natDegree = p := by - simpa only [hp.leadingCoeff, C_1, one_mul, eraseLead] using eraseLead_add_C_mul_X_pow p - by_cases h0 : erase p.natDegree p = 0 - · rw [← H, h0, zero_add] - refine ⟨X ^ p.natDegree, ?_, ?_, monic_X_pow p.natDegree⟩ - · rw [Polynomial.map_pow, map_X] - · rw [degree_X_pow, degree_X_pow] - obtain ⟨q, hq⟩ := mem_lifts_and_degree_eq (erase_mem_lifts p.natDegree hlifts) - have p_neq_0 : p ≠ 0 := by intro hp; apply h0; rw [hp]; simp only [natDegree_zero, erase_zero] - have hdeg : q.degree < ((X : R[X]) ^ p.natDegree).degree := by - rw [@degree_X_pow R, hq.2, ← degree_eq_natDegree p_neq_0] - exact degree_erase_lt p_neq_0 - refine ⟨q + X ^ p.natDegree, ?_, ?_, (monic_X_pow _).add_of_right hdeg⟩ - · rw [Polynomial.map_add, hq.1, Polynomial.map_pow, map_X, H] - · rw [degree_add_eq_right_of_degree_lt hdeg, degree_X_pow, degree_eq_natDegree hp.ne_zero] + rw [lifts_iff_coeff_lifts] at hlifts + let g : ℕ → R := fun k ↦ (hlifts k).choose + have hg k : f (g k) = p.coeff k := (hlifts k).choose_spec + let q : R[X] := X ^ p.natDegree + ∑ k ∈ Finset.range p.natDegree, C (g k) * X ^ k + have hq : map f q = p := by + simp_rw [q, Polynomial.map_add, Polynomial.map_sum, Polynomial.map_mul, Polynomial.map_pow, + map_X, map_C, hg, ← hp.as_sum] + have h : q.Monic := monic_X_pow_add (by simp_rw [← Fin.sum_univ_eq_sum_range, degree_sum_fin_lt]) + exact ⟨q, hq, hq ▸ (h.degree_map f).symm, h⟩ theorem lifts_and_natDegree_eq_and_monic {p : S[X]} (hlifts : p ∈ lifts f) (hp : p.Monic) : ∃ q : R[X], map f q = p ∧ q.natDegree = p.natDegree ∧ q.Monic := by From b10de21dc94402064b24e3fcd210ff4217c7712e Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Sun, 17 Nov 2024 17:51:16 +0000 Subject: [PATCH 037/829] feat(AlgebraicGeometry): basic API for `Proj` (#19152) --- Mathlib.lean | 1 + Mathlib/AlgebraicGeometry/OpenImmersion.lean | 12 + .../ProjectiveSpectrum/Basic.lean | 212 ++++++++++++++++++ .../HomogeneousLocalization.lean | 9 + 4 files changed, 234 insertions(+) create mode 100644 Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean diff --git a/Mathlib.lean b/Mathlib.lean index 928ce53badc19..4210e60788d36 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -960,6 +960,7 @@ import Mathlib.AlgebraicGeometry.PrimeSpectrum.Maximal import Mathlib.AlgebraicGeometry.PrimeSpectrum.Module import Mathlib.AlgebraicGeometry.PrimeSpectrum.Noetherian import Mathlib.AlgebraicGeometry.PrimeSpectrum.TensorProduct +import Mathlib.AlgebraicGeometry.ProjectiveSpectrum.Basic import Mathlib.AlgebraicGeometry.ProjectiveSpectrum.Scheme import Mathlib.AlgebraicGeometry.ProjectiveSpectrum.StructureSheaf import Mathlib.AlgebraicGeometry.ProjectiveSpectrum.Topology diff --git a/Mathlib/AlgebraicGeometry/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/OpenImmersion.lean index ce13477fdd079..de7842550570e 100644 --- a/Mathlib/AlgebraicGeometry/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/OpenImmersion.lean @@ -104,6 +104,18 @@ lemma image_top_eq_opensRange : f ''ᵁ ⊤ = f.opensRange := by apply Opens.ext simp +lemma opensRange_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) + [IsOpenImmersion f] [IsOpenImmersion g] : (f ≫ g).opensRange = g ''ᵁ f.opensRange := + TopologicalSpace.Opens.ext (Set.range_comp g.base f.base) + +lemma opensRange_of_isIso {X Y : Scheme} (f : X ⟶ Y) [IsIso f] : + f.opensRange = ⊤ := + TopologicalSpace.Opens.ext (Set.range_eq_univ.mpr f.homeomorph.surjective) + +lemma opensRange_comp_of_isIso {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) + [IsIso f] [IsOpenImmersion g] : (f ≫ g).opensRange = g.opensRange := by + rw [opensRange_comp, opensRange_of_isIso, image_top_eq_opensRange] + @[simp] lemma preimage_image_eq (U : X.Opens) : f ⁻¹ᵁ f ''ᵁ U = U := by apply Opens.ext diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean new file mode 100644 index 0000000000000..b29a137f71eb2 --- /dev/null +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean @@ -0,0 +1,212 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.AlgebraicGeometry.ProjectiveSpectrum.Scheme +import Mathlib.AlgebraicGeometry.AffineScheme + +/-! + +# Basic properties of the scheme `Proj A` + +The scheme `Proj 𝒜` for a graded algebra `𝒜` is constructed in +`AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean`. +In this file we provide basic properties of the scheme. + +## Main results +- `AlgebraicGeometry.Proj.toSpecZero`: The structure map `Proj A ⟶ Spec (A 0)`. +- `AlgebraicGeometry.Proj.basicOpenIsoSpec`: + The canonical isomorphism `Proj A |_ D₊(f) ≅ Spec (A_f)₀` + when `f` is homogeneous of positive degree. +- `AlgebraicGeometry.Proj.awayι`: The open immersion `Spec (A_f)₀ ⟶ Proj A`. +- `AlgebraicGeometry.Proj.affineOpenCover`: The open cover of `Proj A` by `Spec (A_f)₀` for all + homogeneous `f` of positive degree. +- `AlgebraicGeometry.Proj.stalkIso`: +The stalk of `Proj A` at `x` is the degree `0` part of the localization of `A` at `x`. + +-/ + +namespace AlgebraicGeometry.Proj + +open HomogeneousLocalization CategoryTheory + +variable {R A : Type*} +variable [CommRing R] [CommRing A] [Algebra R A] +variable (𝒜 : ℕ → Submodule R A) +variable [GradedAlgebra 𝒜] + +section basicOpen + +variable (f g : A) + +/-- The basic open set `D₊(f)` associated to `f : A`. -/ +def basicOpen : (Proj 𝒜).Opens := + ProjectiveSpectrum.basicOpen 𝒜 f + +@[simp] +theorem mem_basicOpen (x : Proj 𝒜) : + x ∈ basicOpen 𝒜 f ↔ f ∉ x.asHomogeneousIdeal := + Iff.rfl + +@[simp] theorem basicOpen_one : basicOpen 𝒜 1 = ⊤ := ProjectiveSpectrum.basicOpen_one .. + +@[simp] theorem basicOpen_zero : basicOpen 𝒜 0 = ⊥ := ProjectiveSpectrum.basicOpen_zero .. + +@[simp] theorem basicOpen_pow (n) (hn : 0 < n) : basicOpen 𝒜 (f ^ n) = basicOpen 𝒜 f := + ProjectiveSpectrum.basicOpen_pow 𝒜 f n hn + +theorem basicOpen_mul : basicOpen 𝒜 (f * g) = basicOpen 𝒜 f ⊓ basicOpen 𝒜 g := + ProjectiveSpectrum.basicOpen_mul .. + +theorem basicOpen_mono (hfg : f ∣ g) : basicOpen 𝒜 g ≤ basicOpen 𝒜 f := + (hfg.choose_spec ▸ basicOpen_mul 𝒜 f _).trans_le inf_le_left + +theorem basicOpen_eq_iSup_proj (f : A) : + basicOpen 𝒜 f = ⨆ i : ℕ, basicOpen 𝒜 (GradedAlgebra.proj 𝒜 i f) := + ProjectiveSpectrum.basicOpen_eq_union_of_projection .. + +theorem isBasis_basicOpen : + TopologicalSpace.Opens.IsBasis (Set.range (basicOpen 𝒜)) := by + delta TopologicalSpace.Opens.IsBasis + convert ProjectiveSpectrum.isTopologicalBasis_basic_opens 𝒜 + exact (Set.range_comp _ _).symm + +lemma iSup_basicOpen_eq_top {ι : Type*} (f : ι → A) + (hf : (HomogeneousIdeal.irrelevant 𝒜).toIdeal ≤ Ideal.span (Set.range f)) : + ⨆ i, Proj.basicOpen 𝒜 (f i) = ⊤ := by + classical + refine top_le_iff.mp fun x hx ↦ TopologicalSpace.Opens.mem_iSup.mpr ?_ + by_contra! H + simp only [mem_basicOpen, Decidable.not_not] at H + refine x.not_irrelevant_le (hf.trans ?_) + rwa [Ideal.span_le, Set.range_subset_iff] + +/-- The canonical map `(A_f)₀ ⟶ Γ(Proj A, D₊(f))`. +This is an isomorphism when `f` is homogeneous of positive degree. See `basicOpenIsoAway` below. -/ +def awayToSection : CommRingCat.of (Away 𝒜 f) ⟶ Γ(Proj 𝒜, basicOpen 𝒜 f) := + ProjectiveSpectrum.Proj.awayToSection .. + +/-- The canonical map `Proj A |_ D₊(f) ⟶ Spec (A_f)₀`. +This is an isomorphism when `f` is homogeneous of positive degree. See `basicOpenIsoSpec` below. -/ +noncomputable +def basicOpenToSpec : (basicOpen 𝒜 f).toScheme ⟶ Spec (.of (Away 𝒜 f)) := + (basicOpen 𝒜 f).toSpecΓ ≫ Spec.map (awayToSection 𝒜 f) + +lemma basicOpenToSpec_app_top : + (basicOpenToSpec 𝒜 f).app ⊤ = (Scheme.ΓSpecIso _).hom ≫ awayToSection 𝒜 f ≫ + (basicOpen 𝒜 f).topIso.inv := by + rw [basicOpenToSpec] + simp only [Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, + TopologicalSpace.Opens.map_top, Scheme.comp_app, Scheme.Opens.topIso_inv, eqToHom_op] + erw [Scheme.comp_app] + simp + +/-- The structure map `Proj A ⟶ Spec A₀`. -/ +noncomputable +def toSpecZero : Proj 𝒜 ⟶ Spec (.of (𝒜 0)) := + (Scheme.topIso _).inv ≫ (Scheme.isoOfEq _ (basicOpen_one _)).inv ≫ + basicOpenToSpec 𝒜 1 ≫ Spec.map (fromZeroRingHom 𝒜 _) + +variable {m} (f_deg : f ∈ 𝒜 m) (hm : 0 < m) + +/-- The canonical isomorphism `Proj A |_ D₊(f) ≅ Spec (A_f)₀` +when `f` is homogeneous of positive degree. -/ +@[simps! (config := .lemmasOnly) hom] +noncomputable +def basicOpenIsoSpec : (basicOpen 𝒜 f).toScheme ≅ Spec (.of (Away 𝒜 f)) := + have : IsIso (basicOpenToSpec 𝒜 f) := by + apply (isIso_iff_of_reflects_iso _ Scheme.forgetToLocallyRingedSpace).mp ?_ + convert ProjectiveSpectrum.Proj.isIso_toSpec 𝒜 f f_deg hm using 1 + refine Eq.trans ?_ (ΓSpec.locallyRingedSpaceAdjunction.homEquiv_apply _ _ _).symm + dsimp [basicOpenToSpec, Scheme.Opens.toSpecΓ] + simp only [eqToHom_op, Category.assoc, ← Spec.map_comp] + rfl + asIso (basicOpenToSpec 𝒜 f) + +/-- The canonical isomorphism `(A_f)₀ ≅ Γ(Proj A, D₊(f))` +when `f` is homogeneous of positive degree. -/ +@[simps! (config := .lemmasOnly) hom] +noncomputable +def basicOpenIsoAway : CommRingCat.of (Away 𝒜 f) ≅ Γ(Proj 𝒜, basicOpen 𝒜 f) := + have : IsIso (awayToSection 𝒜 f) := by + have := basicOpenToSpec_app_top 𝒜 f + rw [← Iso.inv_comp_eq, Iso.eq_comp_inv] at this + rw [← this, ← basicOpenIsoSpec_hom 𝒜 f f_deg hm] + infer_instance + asIso (awayToSection 𝒜 f) + +/-- The open immersion `Spec (A_f)₀ ⟶ Proj A`. -/ +noncomputable +def awayι : Spec (.of (Away 𝒜 f)) ⟶ Proj 𝒜 := + (basicOpenIsoSpec 𝒜 f f_deg hm).inv ≫ (Proj.basicOpen 𝒜 f).ι + +instance : IsOpenImmersion (Proj.awayι 𝒜 f f_deg hm) := + IsOpenImmersion.comp _ _ + +lemma opensRange_awayι : + (Proj.awayι 𝒜 f f_deg hm).opensRange = Proj.basicOpen 𝒜 f := + (Scheme.Hom.opensRange_comp_of_isIso _ _).trans (basicOpen 𝒜 f).opensRange_ι + +include f_deg hm in +lemma isAffineOpen_basicOpen : IsAffineOpen (basicOpen 𝒜 f) := by + rw [← opensRange_awayι 𝒜 f f_deg hm] + exact isAffineOpen_opensRange (awayι _ _ _ _) + +@[reassoc] +lemma awayι_toSpecZero : awayι 𝒜 f f_deg hm ≫ toSpecZero 𝒜 = + Spec.map (CommRingCat.ofHom (fromZeroRingHom 𝒜 _)) := by + rw [toSpecZero, basicOpenToSpec, awayι] + simp only [Category.assoc, Iso.inv_comp_eq, basicOpenIsoSpec_hom] + have (U) (e : U = ⊤) : (basicOpen 𝒜 f).ι ≫ (Scheme.topIso _).inv ≫ (Scheme.isoOfEq _ e).inv = + Scheme.homOfLE _ (le_top.trans_eq e.symm) := by + simp only [← Category.assoc, Iso.comp_inv_eq] + simp only [Scheme.topIso_hom, Category.assoc, Scheme.isoOfEq_hom_ι, Scheme.homOfLE_ι] + rw [reassoc_of% this, ← Scheme.Opens.toSpecΓ_SpecMap_map_assoc, basicOpenToSpec, Category.assoc, + ← Spec.map_comp, ← Spec.map_comp, ← Spec.map_comp] + rfl + +open TopologicalSpace.Opens in +/-- Given a family of homogeneous elements `f` of positive degree that spans the irrelavent ideal, +`Spec (A_f)₀ ⟶ Proj A` forms an affine open cover of `Proj A`. -/ +noncomputable +def openCoverOfISupEqTop {ι : Type*} (f : ι → A) {m : ι → ℕ} + (f_deg : ∀ i, f i ∈ 𝒜 (m i)) (hm : ∀ i, 0 < m i) + (hf : (HomogeneousIdeal.irrelevant 𝒜).toIdeal ≤ Ideal.span (Set.range f)) : + (Proj 𝒜).AffineOpenCover where + J := ι + obj i := .of (Away 𝒜 (f i)) + map i := awayι 𝒜 (f i) (f_deg i) (hm i) + f x := (mem_iSup.mp ((iSup_basicOpen_eq_top 𝒜 f hf).ge (Set.mem_univ x))).choose + covers x := by + show x ∈ (awayι 𝒜 _ _ _).opensRange + rw [opensRange_awayι] + exact (mem_iSup.mp ((iSup_basicOpen_eq_top 𝒜 f hf).ge (Set.mem_univ x))).choose_spec + +/-- `Proj A` is covered by `Spec (A_f)₀` for all homogeneous elements of positive degree. -/ +noncomputable +def affineOpenCover : (Proj 𝒜).AffineOpenCover := + openCoverOfISupEqTop 𝒜 (ι := Σ i : PNat, 𝒜 i) (m := fun i ↦ i.1) (fun i ↦ i.2) (fun i ↦ i.2.2) + (fun i ↦ i.1.2) <| by + classical + intro z hz + rw [← DirectSum.sum_support_decompose 𝒜 z] + refine Ideal.sum_mem _ fun c hc ↦ if hc0 : c = 0 then ?_ else + Ideal.subset_span ⟨⟨⟨c, Nat.pos_iff_ne_zero.mpr hc0⟩, _⟩, rfl⟩ + convert Ideal.zero_mem _ + subst hc0 + exact hz + +end basicOpen + +section stalk + +/-- The stalk of `Proj A` at `x` is the degree `0` part of the localization of `A` at `x`. -/ +noncomputable +def stalkIso (x : Proj 𝒜) : + (Proj 𝒜).presheaf.stalk x ≅ .of (AtPrime 𝒜 x.asHomogeneousIdeal.toIdeal) := + (stalkIso' 𝒜 x).toCommRingCatIso + +end stalk + +end AlgebraicGeometry.Proj diff --git a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean index 8eb554c57e109..553951dd2ce18 100644 --- a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean +++ b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean @@ -464,6 +464,15 @@ lemma mk_eq_zero_of_den (f : NumDenSameDeg 𝒜 x) (h : f.den = 0) : mk f = 0 := have := subsingleton 𝒜 (h ▸ f.den_mem) exact Subsingleton.elim _ _ +variable (𝒜 x) in +/-- The map from `𝒜 0` to the degree `0` part of `𝒜ₓ` sending `f ↦ f/1`. -/ +def fromZeroRingHom : 𝒜 0 →+* HomogeneousLocalization 𝒜 x where + toFun f := .mk ⟨0, f, 1, one_mem _⟩ + map_one' := rfl + map_mul' f g := by ext; simp [Localization.mk_mul] + map_zero' := rfl + map_add' f g := by ext; simp [Localization.add_mk, add_comm f.1 g.1] + end HomogeneousLocalization namespace HomogeneousLocalization From 391a4497fae3e05240d9cfe682ffbdd85ae7481d Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Sun, 17 Nov 2024 18:08:40 +0000 Subject: [PATCH 038/829] feat(Algebra/Group/Aut): Isomorphic groups have isomorphic automorphism groups (#19007) This PR shows that isomorphic groups have isomorphic automorphism groups. --- Mathlib/Algebra/Group/Aut.lean | 20 ++++++++++++++++++++ 1 file changed, 20 insertions(+) diff --git a/Mathlib/Algebra/Group/Aut.lean b/Mathlib/Algebra/Group/Aut.lean index 567b51be95470..ac678be0258c9 100644 --- a/Mathlib/Algebra/Group/Aut.lean +++ b/Mathlib/Algebra/Group/Aut.lean @@ -141,6 +141,16 @@ theorem conj_symm_apply [Group G] (g h : G) : (conj g).symm h = g⁻¹ * h * g : theorem conj_inv_apply [Group G] (g h : G) : (conj g)⁻¹ h = g⁻¹ * h * g := rfl +/-- Isomorphic groups have isomorphic automorphism groups. -/ +@[simps] +def congr [Group G] {H : Type*} [Group H] (ϕ : G ≃* H) : + MulAut G ≃* MulAut H where + toFun f := ϕ.symm.trans (f.trans ϕ) + invFun f := ϕ.trans (f.trans ϕ.symm) + left_inv _ := by simp [DFunLike.ext_iff] + right_inv _ := by simp [DFunLike.ext_iff] + map_mul' := by simp [DFunLike.ext_iff] + end MulAut namespace AddAut @@ -255,4 +265,14 @@ theorem conj_symm_apply [AddGroup G] (g h : G) : (conj g).symm h = -g + h + g := theorem conj_inv_apply [AddGroup G] (g h : G) : (Additive.toMul (conj g))⁻¹ h = -g + h + g := rfl +/-- Isomorphic additive groups have isomorphic automorphism groups. -/ +@[simps] +def congr [AddGroup G] {H : Type*} [AddGroup H] (ϕ : G ≃+ H) : + AddAut G ≃* AddAut H where + toFun f := ϕ.symm.trans (f.trans ϕ) + invFun f := ϕ.trans (f.trans ϕ.symm) + left_inv _ := by simp [DFunLike.ext_iff] + right_inv _ := by simp [DFunLike.ext_iff] + map_mul' := by simp [DFunLike.ext_iff] + end AddAut From aa0a9db5a44caf6ab1d2f53737ba346ab2c0f4b5 Mon Sep 17 00:00:00 2001 From: Julian Berman Date: Sun, 17 Nov 2024 18:32:46 +0000 Subject: [PATCH 039/829] chore: upload import graphs on build (#19162) Previously we were building them then throwing them away, but if we have them already it seems nice to expose them for use. [Zulip](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Uploading.20the.20import.20graph) --- .github/build.in.yml | 18 ++++++++++++++---- .github/workflows/bors.yml | 18 ++++++++++++++---- .github/workflows/build.yml | 18 ++++++++++++++---- .github/workflows/build_fork.yml | 18 ++++++++++++++---- 4 files changed, 56 insertions(+), 16 deletions(-) diff --git a/.github/build.in.yml b/.github/build.in.yml index 1b4a95610f10a..3afe6166a83f4 100644 --- a/.github/build.in.yml +++ b/.github/build.in.yml @@ -240,10 +240,20 @@ jobs: python3 scripts/yaml_check.py docs/100.yaml docs/overview.yaml docs/undergrad.yaml lake exe check-yaml - - name: verify `lake exe graph` works - run: | - lake exe graph - rm import_graph.dot + - name: generate our import graph + run: lake exe graph + + - name: upload the import graph + uses: actions/upload-artifact@v4 + with: + name: import-graph + path: import_graph.dot + ## the default is 90, but we build often, so unless there's a reason + ## to care about old copies in the future, just say 7 days for now + retention-days: 7 + + - name: clean up the import graph file + run: rm import_graph.dot - name: build everything # make sure everything is available for test/import_all.lean diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index 114e041faea32..598668edd66d3 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -250,10 +250,20 @@ jobs: python3 scripts/yaml_check.py docs/100.yaml docs/overview.yaml docs/undergrad.yaml lake exe check-yaml - - name: verify `lake exe graph` works - run: | - lake exe graph - rm import_graph.dot + - name: generate our import graph + run: lake exe graph + + - name: upload the import graph + uses: actions/upload-artifact@v4 + with: + name: import-graph + path: import_graph.dot + ## the default is 90, but we build often, so unless there's a reason + ## to care about old copies in the future, just say 7 days for now + retention-days: 7 + + - name: clean up the import graph file + run: rm import_graph.dot - name: build everything # make sure everything is available for test/import_all.lean diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 4772983e9c353..4afc15b2f5939 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -257,10 +257,20 @@ jobs: python3 scripts/yaml_check.py docs/100.yaml docs/overview.yaml docs/undergrad.yaml lake exe check-yaml - - name: verify `lake exe graph` works - run: | - lake exe graph - rm import_graph.dot + - name: generate our import graph + run: lake exe graph + + - name: upload the import graph + uses: actions/upload-artifact@v4 + with: + name: import-graph + path: import_graph.dot + ## the default is 90, but we build often, so unless there's a reason + ## to care about old copies in the future, just say 7 days for now + retention-days: 7 + + - name: clean up the import graph file + run: rm import_graph.dot - name: build everything # make sure everything is available for test/import_all.lean diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index 90535d997a9d6..1964290d30bc7 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -254,10 +254,20 @@ jobs: python3 scripts/yaml_check.py docs/100.yaml docs/overview.yaml docs/undergrad.yaml lake exe check-yaml - - name: verify `lake exe graph` works - run: | - lake exe graph - rm import_graph.dot + - name: generate our import graph + run: lake exe graph + + - name: upload the import graph + uses: actions/upload-artifact@v4 + with: + name: import-graph + path: import_graph.dot + ## the default is 90, but we build often, so unless there's a reason + ## to care about old copies in the future, just say 7 days for now + retention-days: 7 + + - name: clean up the import graph file + run: rm import_graph.dot - name: build everything # make sure everything is available for test/import_all.lean From 08ac4589995d198221a35f21c83467aba483fb3a Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Sun, 17 Nov 2024 18:41:57 +0000 Subject: [PATCH 040/829] feat: towards equalizers in the category of ind-objects (#18522) We're still quite far away from the final goal `ClosedUnderLimitsOfShape WalkingParallelPair (IsIndObject (C := C))`, but this is some progress in that direction. Co-authored-by: Markus Himmel --- Mathlib.lean | 1 + .../Limits/Indization/Equalizers.lean | 63 +++++++++++++++++++ 2 files changed, 64 insertions(+) create mode 100644 Mathlib/CategoryTheory/Limits/Indization/Equalizers.lean diff --git a/Mathlib.lean b/Mathlib.lean index 4210e60788d36..3b1c7dbea4b29 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1777,6 +1777,7 @@ import Mathlib.CategoryTheory.Limits.FunctorToTypes import Mathlib.CategoryTheory.Limits.HasLimits import Mathlib.CategoryTheory.Limits.IndYoneda import Mathlib.CategoryTheory.Limits.Indization.Category +import Mathlib.CategoryTheory.Limits.Indization.Equalizers import Mathlib.CategoryTheory.Limits.Indization.FilteredColimits import Mathlib.CategoryTheory.Limits.Indization.IndObject import Mathlib.CategoryTheory.Limits.Indization.LocallySmall diff --git a/Mathlib/CategoryTheory/Limits/Indization/Equalizers.lean b/Mathlib/CategoryTheory/Limits/Indization/Equalizers.lean new file mode 100644 index 0000000000000..ef9fd1afa688d --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Indization/Equalizers.lean @@ -0,0 +1,63 @@ +/- +Copyright (c) 2024 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel +-/ +import Mathlib.CategoryTheory.Limits.Indization.FilteredColimits + +/-! +# Equalizers of ind-objects + +The eventual goal of this file is to show that if a category `C` has equalizers, then ind-objects +are closed under equalizers. For now, it contains one of the main prerequisites, namely we show +that under sufficient conditions the limit of a diagram in `I ⥤ C` taken in `Cᵒᵖ ⥤ Type v` is an +ind-object. + +## References +* [M. Kashiwara, P. Schapira, *Categories and Sheaves*][Kashiwara2006], Section 6.1 +-/ + +universe v v' u u' + +namespace CategoryTheory.Limits + +variable {C : Type u} [Category.{v} C] + +section + +variable {I : Type v} [SmallCategory I] [IsFiltered I] + +variable {J : Type} [SmallCategory J] [FinCategory J] + +variable (F : J ⥤ I ⥤ C) + +/-- +Suppose `F : J ⥤ I ⥤ C` is a finite diagram in the functor category `I ⥤ C`, where `I` is small +and filtered. If `i : I`, we can apply the Yoneda embedding to `F(·, i)` to obtain a +diagram of presheaves `J ⥤ Cᵒᵖ ⥤ Type v`. Suppose that the limits of this diagram is always an +ind-object. + +For `j : J` we can apply the Yoneda embedding to `F(j, ·)` and take colimits to obtain a finite +diagram `J ⥤ Cᵒᵖ ⥤ Type v` (which is actually a diagram `J ⥤ Ind C`). The theorem states that +the limit of this diagram is an ind-object. + +This theorem will be used to construct equalizers in the category of ind-objects. It can be +interpreted as saying that ind-objects are closed under finite limits as long as the diagram +we are taking the limit of comes from a diagram in a functor category `I ⥤ C`. We will show (TODO) +that this is the case for any parallel pair of morphisms in `Ind C` and deduce that ind-objects +are closed under equalizers. + +This is Proposition 6.1.16(i) in [Kashiwara2006]. +-/ +theorem isIndObject_limit_comp_yoneda_comp_colim + (hF : ∀ i, IsIndObject (limit (F.flip.obj i ⋙ yoneda))) : + IsIndObject (limit (F ⋙ (whiskeringRight _ _ _).obj yoneda ⋙ colim)) := by + let G : J ⥤ I ⥤ (Cᵒᵖ ⥤ Type v) := F ⋙ (whiskeringRight _ _ _).obj yoneda + apply IsIndObject.map (HasLimit.isoOfNatIso (colimitFlipIsoCompColim G)).hom + apply IsIndObject.map (colimitLimitIso G).hom + apply isIndObject_colimit + exact fun i => IsIndObject.map (limitObjIsoLimitCompEvaluation _ _).inv (hF i) + +end + +end CategoryTheory.Limits From 38affacc35abeefc7fd0c782d547d6b745281ad7 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Sun, 17 Nov 2024 19:21:38 +0000 Subject: [PATCH 041/829] chore: use AlgHom.restrictDomain (#19116) I previously didn't know the existence of `AlgHom.restrictDomain` and wrote many lemmas using a longer expression that is exactly the definition of `AlgHom.restrictDomain`; I'm now changing all these lemmas to use `AlgHom.restrictDomain`. There are more appearances of `.comp (IsScalarTower.toAlgHom` but either the file doesn't import `AlgHom.restrictDomain`, or using it breaks the proof, so I only replaced one appearance. --- Mathlib/FieldTheory/Extension.lean | 11 +++++------ Mathlib/FieldTheory/IsAlgClosed/Basic.lean | 7 +++++-- Mathlib/FieldTheory/IsSepClosed.lean | 7 +++++-- .../RingTheory/IntegralClosure/IntegralRestrict.lean | 12 ++++++------ Mathlib/RingTheory/Unramified/Basic.lean | 2 +- 5 files changed, 22 insertions(+), 17 deletions(-) diff --git a/Mathlib/FieldTheory/Extension.lean b/Mathlib/FieldTheory/Extension.lean index abfd3dbf451d7..da3e16176910d 100644 --- a/Mathlib/FieldTheory/Extension.lean +++ b/Mathlib/FieldTheory/Extension.lean @@ -96,7 +96,7 @@ section private theorem exists_algHom_adjoin_of_splits'' {L : IntermediateField F E} (f : L →ₐ[F] K) (hK : ∀ s ∈ S, IsIntegral L s ∧ (minpoly L s).Splits f.toRingHom) : - ∃ φ : adjoin L S →ₐ[F] K, φ.comp (IsScalarTower.toAlgHom F L _) = f := by + ∃ φ : adjoin L S →ₐ[F] K, φ.restrictDomain L = f := by obtain ⟨φ, hfφ, hφ⟩ := zorn_le_nonempty_Ici₀ _ (fun c _ hc _ _ ↦ Lifts.exists_upper_bound c hc) ⟨L, f⟩ le_rfl refine ⟨φ.emb.comp (inclusion <| (le_extendScalars_iff hfφ.1 <| adjoin L S).mp <| @@ -114,7 +114,7 @@ variable {L : Type*} [Field L] [Algebra F L] [Algebra L E] [IsScalarTower F L E] include hK in theorem exists_algHom_adjoin_of_splits' : - ∃ φ : adjoin L S →ₐ[F] K, φ.comp (IsScalarTower.toAlgHom F L _) = f := by + ∃ φ : adjoin L S →ₐ[F] K, φ.restrictDomain L = f := by let L' := (IsScalarTower.toAlgHom F L E).fieldRange let f' : L' →ₐ[F] K := f.comp (AlgEquiv.ofInjectiveField _).symm.toAlgHom have := exists_algHom_adjoin_of_splits'' f' (S := S) fun s hs ↦ ?_ @@ -123,21 +123,20 @@ theorem exists_algHom_adjoin_of_splits' : · simp_rw [← SetLike.coe_subset_coe, coe_restrictScalars, adjoin_subset_adjoin_iff] exact ⟨subset_adjoin_of_subset_left S (F := L'.toSubfield) le_rfl, subset_adjoin _ _⟩ · ext x - rw [AlgHom.comp_assoc] exact congr($hφ _).trans (congr_arg f <| AlgEquiv.symm_apply_apply _ _) - letI : Algebra L L' := (AlgEquiv.ofInjectiveField _).toRingEquiv.toRingHom.toAlgebra + letI : Algebra L L' := (AlgEquiv.ofInjectiveField _).toRingHom.toAlgebra have : IsScalarTower L L' E := IsScalarTower.of_algebraMap_eq' rfl refine ⟨(hK s hs).1.tower_top, (hK s hs).1.minpoly_splits_tower_top' ?_⟩ convert (hK s hs).2; ext; exact congr_arg f (AlgEquiv.symm_apply_apply _ _) include hK in theorem exists_algHom_of_adjoin_splits' (hS : adjoin L S = ⊤) : - ∃ φ : E →ₐ[F] K, φ.comp (IsScalarTower.toAlgHom F L E) = f := + ∃ φ : E →ₐ[F] K, φ.restrictDomain L = f := have ⟨φ, hφ⟩ := exists_algHom_adjoin_of_splits' f hK ⟨φ.comp (((equivOfEq hS).trans topEquiv).symm.toAlgHom.restrictScalars F), hφ⟩ theorem exists_algHom_of_splits' (hK : ∀ s : E, IsIntegral L s ∧ (minpoly L s).Splits f.toRingHom) : - ∃ φ : E →ₐ[F] K, φ.comp (IsScalarTower.toAlgHom F L E) = f := + ∃ φ : E →ₐ[F] K, φ.restrictDomain L = f := exists_algHom_of_adjoin_splits' f (fun x _ ↦ hK x) (adjoin_univ L E) end diff --git a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean index b852e2ebe37f6..2ec4275183cac 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean @@ -262,12 +262,15 @@ variable {K : Type u} [Field K] {L : Type v} {M : Type w} [Field L] [Algebra K L /-- If E/L/K is a tower of field extensions with E/L algebraic, and if M is an algebraically closed extension of K, then any embedding of L/K into M/K extends to an embedding of E/K. Known as the extension lemma in https://math.stackexchange.com/a/687914. -/ -theorem surjective_comp_algebraMap_of_isAlgebraic {E : Type*} +theorem surjective_restrictDomain_of_isAlgebraic {E : Type*} [Field E] [Algebra K E] [Algebra L E] [IsScalarTower K L E] [Algebra.IsAlgebraic L E] : - Function.Surjective fun φ : E →ₐ[K] M ↦ φ.comp (IsScalarTower.toAlgHom K L E) := + Function.Surjective fun φ : E →ₐ[K] M ↦ φ.restrictDomain L := fun f ↦ IntermediateField.exists_algHom_of_splits' (E := E) f fun s ↦ ⟨Algebra.IsIntegral.isIntegral s, IsAlgClosed.splits_codomain _⟩ +@[deprecated (since := "2024-11-15")] +alias surjective_comp_algebraMap_of_isAlgebraic := surjective_restrictDomain_of_isAlgebraic + variable [Algebra.IsAlgebraic K L] (K L M) /-- Less general version of `lift`. -/ diff --git a/Mathlib/FieldTheory/IsSepClosed.lean b/Mathlib/FieldTheory/IsSepClosed.lean index 3de1da530b30c..e3f8863e5ba79 100644 --- a/Mathlib/FieldTheory/IsSepClosed.lean +++ b/Mathlib/FieldTheory/IsSepClosed.lean @@ -269,13 +269,16 @@ namespace IsSepClosed variable {K : Type u} (L : Type v) {M : Type w} [Field K] [Field L] [Algebra K L] [Field M] [Algebra K M] [IsSepClosed M] -theorem surjective_comp_algebraMap_of_isSeparable {E : Type*} +theorem surjective_restrictDomain_of_isSeparable {E : Type*} [Field E] [Algebra K E] [Algebra L E] [IsScalarTower K L E] [Algebra.IsSeparable L E] : - Function.Surjective fun φ : E →ₐ[K] M ↦ φ.comp (IsScalarTower.toAlgHom K L E) := + Function.Surjective fun φ : E →ₐ[K] M ↦ φ.restrictDomain L := fun f ↦ IntermediateField.exists_algHom_of_splits' (E := E) f fun s ↦ ⟨Algebra.IsSeparable.isIntegral L s, IsSepClosed.splits_codomain _ <| Algebra.IsSeparable.isSeparable L s⟩ +@[deprecated (since := "2024-11-15")] +alias surjective_comp_algebraMap_of_isSeparable := surjective_restrictDomain_of_isSeparable + variable [Algebra.IsSeparable K L] {L} /-- A (random) homomorphism from a separable extension L of K into a separably diff --git a/Mathlib/RingTheory/IntegralClosure/IntegralRestrict.lean b/Mathlib/RingTheory/IntegralClosure/IntegralRestrict.lean index 5253fd9dff111..1c84755bf4e23 100644 --- a/Mathlib/RingTheory/IntegralClosure/IntegralRestrict.lean +++ b/Mathlib/RingTheory/IntegralClosure/IntegralRestrict.lean @@ -66,8 +66,8 @@ def galRestrictHom : (L →ₐ[K] L) ≃* (B →ₐ[A] B) where apply (IsIntegralClosure.equiv A (integralClosure A L) L B).symm.injective ext dsimp - simp only [AlgEquiv.symm_apply_apply, AlgHom.coe_codRestrict, AlgHom.coe_comp, - AlgHom.coe_restrictScalars', IsScalarTower.coe_toAlgHom', Function.comp_apply, + simp only [AlgEquiv.symm_apply_apply, AlgHom.coe_codRestrict, AlgHom.coe_restrictScalars', + AlgHom.coe_comp, AlgHom.restrictDomain, IsScalarTower.coe_toAlgHom', Function.comp_apply, AlgHom.mul_apply, IsIntegralClosure.algebraMap_equiv, Subalgebra.algebraMap_eq] rfl invFun := galLift A K L B @@ -75,17 +75,17 @@ def galRestrictHom : (L →ₐ[K] L) ≃* (B →ₐ[A] B) where have := (IsFractionRing.injective A K).isDomain have := IsIntegralClosure.isLocalization A K L B AlgHom.coe_ringHom_injective <| IsLocalization.ringHom_ext (Algebra.algebraMapSubmonoid B A⁰) - <| RingHom.ext fun x ↦ by simp [Subalgebra.algebraMap_eq, galLift] + <| RingHom.ext fun x ↦ by simp [Subalgebra.algebraMap_eq, AlgHom.restrictDomain, galLift] right_inv σ := have := (IsFractionRing.injective A K).isDomain have := IsIntegralClosure.isLocalization A K L B - AlgHom.ext fun x ↦ - IsIntegralClosure.algebraMap_injective B A L (by simp [Subalgebra.algebraMap_eq, galLift]) + AlgHom.ext fun x ↦ IsIntegralClosure.algebraMap_injective B A L + (by simp [AlgHom.restrictDomain, Subalgebra.algebraMap_eq, galLift]) @[simp] lemma algebraMap_galRestrictHom_apply (σ : L →ₐ[K] L) (x : B) : algebraMap B L (galRestrictHom A K L B σ x) = σ (algebraMap B L x) := by - simp [galRestrictHom, Subalgebra.algebraMap_eq] + simp [galRestrictHom, Subalgebra.algebraMap_eq, AlgHom.restrictDomain] @[simp, nolint unusedHavesSuffices] -- false positive from unfolding galRestrictHom lemma galRestrictHom_symm_algebraMap_apply (σ : B →ₐ[A] B) (x : B) : diff --git a/Mathlib/RingTheory/Unramified/Basic.lean b/Mathlib/RingTheory/Unramified/Basic.lean index 73ca2ac3c5f04..5131929b68787 100644 --- a/Mathlib/RingTheory/Unramified/Basic.lean +++ b/Mathlib/RingTheory/Unramified/Basic.lean @@ -174,7 +174,7 @@ theorem comp [FormallyUnramified R A] [FormallyUnramified A B] : have e' := FormallyUnramified.lift_unique I ⟨2, hI⟩ (f₁.comp <| IsScalarTower.toAlgHom R A B) (f₂.comp <| IsScalarTower.toAlgHom R A B) (by rw [← AlgHom.comp_assoc, e, AlgHom.comp_assoc]) - letI := (f₁.comp (IsScalarTower.toAlgHom R A B)).toRingHom.toAlgebra + letI := (f₁.restrictDomain A).toAlgebra let F₁ : B →ₐ[A] C := { f₁ with commutes' := fun r => rfl } let F₂ : B →ₐ[A] C := { f₂ with commutes' := AlgHom.congr_fun e'.symm } ext1 x From 7aeae34dfe1ca991bdb678b042ab1c7e3bea0df7 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Mon, 18 Nov 2024 02:03:40 +0000 Subject: [PATCH 042/829] chore: fix warnings in longest pole (#19077) --- LongestPole/Rectangles.lean | 5 +++-- LongestPole/Unused.lean | 2 +- 2 files changed, 4 insertions(+), 3 deletions(-) diff --git a/LongestPole/Rectangles.lean b/LongestPole/Rectangles.lean index bc1cb181c5269..8855f055de44e 100644 --- a/LongestPole/Rectangles.lean +++ b/LongestPole/Rectangles.lean @@ -70,8 +70,9 @@ def area (r : Rectangle) : Nat := r.width * r.height The number of points in `r`, weighted by the function `w`. -/ def weightedArea (r : Rectangle) (w : Nat × Nat → Nat) : Nat := - Nat.sum <| - (List.range' r.left r.width).bind fun x => (List.range' r.bottom r.height).map fun y => w (x, y) + List.sum <| + (List.range' r.left r.width).flatMap + fun x => (List.range' r.bottom r.height).map fun y => w (x, y) end Rectangle diff --git a/LongestPole/Unused.lean b/LongestPole/Unused.lean index 6633c405110b9..ebf4d74987ec3 100644 --- a/LongestPole/Unused.lean +++ b/LongestPole/Unused.lean @@ -78,7 +78,7 @@ def unusedImportsCLI (args : Cli.Parsed) : IO UInt32 := do IO.println s!"Writing table to {output}." IO.FS.writeFile output (formatTable headings rows.toArray) - let data := unused.bind fun (m, u) => u.map fun n => (modules.indexOf m, modules.indexOf n) + let data := unused.flatMap fun (m, u) => u.map fun n => (modules.indexOf m, modules.indexOf n) let rectangles := maximalRectangles data |>.map (fun r => (r, r.area)) -- Prefer rectangles with larger areas. From 25f427c3569562eff0d0880654a018e7027e7800 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 18 Nov 2024 02:41:21 +0000 Subject: [PATCH 043/829] feat(SetTheory/Cardinal/Aleph): some lemmas on `omega` (#18542) --- Mathlib/SetTheory/Cardinal/Aleph.lean | 23 ++++++++++++++++++++++- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 4 ---- Mathlib/SetTheory/Ordinal/Basic.lean | 3 --- 3 files changed, 22 insertions(+), 8 deletions(-) diff --git a/Mathlib/SetTheory/Cardinal/Aleph.lean b/Mathlib/SetTheory/Cardinal/Aleph.lean index 1212a22dd9926..5d5d56a75f0ee 100644 --- a/Mathlib/SetTheory/Cardinal/Aleph.lean +++ b/Mathlib/SetTheory/Cardinal/Aleph.lean @@ -221,6 +221,10 @@ theorem omega0_le_omega (o : Ordinal) : ω ≤ ω_ o := by rw [← omega_zero, omega_le_omega] exact Ordinal.zero_le o +/-- For the theorem `0 < ω`, see `omega0_pos`. -/ +theorem omega_pos (o : Ordinal) : 0 < ω_ o := + omega0_pos.trans_le (omega0_le_omega o) + theorem omega0_lt_omega1 : ω < ω₁ := by rw [← omega_zero, omega_lt_omega] exact zero_lt_one @@ -306,6 +310,11 @@ theorem lift_preAleph (o : Ordinal.{u}) : lift.{v} (preAleph o) = preAleph (Ordi (preAleph.toInitialSeg.trans liftInitialSeg).eq (Ordinal.liftInitialSeg.trans preAleph.toInitialSeg) o +@[simp] +theorem _root_.Ordinal.lift_preOmega (o : Ordinal.{u}) : + Ordinal.lift.{v} (preOmega o) = preOmega (Ordinal.lift.{v} o) := by + rw [← ord_preAleph, lift_ord, lift_preAleph, ord_preAleph] + theorem preAleph_le_of_isLimit {o : Ordinal} (l : o.IsLimit) {c} : preAleph o ≤ c ↔ ∀ o' < o, preAleph o' ≤ c := ⟨fun h o' h' => (preAleph_le_preAleph.2 <| h'.le).trans h, fun h => by @@ -383,6 +392,12 @@ theorem aleph_zero : ℵ_ 0 = ℵ₀ := by rw [aleph_eq_preAleph, add_zero, preA theorem lift_aleph (o : Ordinal.{u}) : lift.{v} (aleph o) = aleph (Ordinal.lift.{v} o) := by simp [aleph_eq_preAleph] +/-- For the theorem `lift ω = ω`, see `lift_omega0`. -/ +@[simp] +theorem _root_.Ordinal.lift_omega (o : Ordinal.{u}) : + Ordinal.lift.{v} (ω_ o) = ω_ (Ordinal.lift.{v} o) := by + simp [omega_eq_preOmega] + theorem aleph_limit {o : Ordinal} (ho : o.IsLimit) : ℵ_ o = ⨆ a : Iio o, ℵ_ a := by apply le_antisymm _ (ciSup_le' _) · rw [aleph_eq_preAleph, preAleph_limit (ho.add _)] @@ -414,11 +429,17 @@ instance nonempty_toType_aleph (o : Ordinal) : Nonempty (ℵ_ o).ord.toType := b rw [toType_nonempty_iff_ne_zero, ← ord_zero] exact fun h => (ord_injective h).not_gt (aleph_pos o) +theorem isLimit_omega (o : Ordinal) : Ordinal.IsLimit (ω_ o) := by + rw [← ord_aleph] + exact isLimit_ord (aleph0_le_aleph _) + +@[deprecated isLimit_omega (since := "2024-10-24")] theorem ord_aleph_isLimit (o : Ordinal) : (ℵ_ o).ord.IsLimit := isLimit_ord <| aleph0_le_aleph _ +-- TODO: get rid of this instance where it's used. instance (o : Ordinal) : NoMaxOrder (ℵ_ o).ord.toType := - toType_noMax_of_succ_lt (ord_aleph_isLimit o).2 + toType_noMax_of_succ_lt (isLimit_ord <| aleph0_le_aleph o).2 @[simp] theorem range_aleph : range aleph = Set.Ici ℵ₀ := by diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 650de1e4c9277..6cc7dfe51df9e 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -2304,10 +2304,6 @@ 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_omega0 0 - theorem omega0_ne_zero : ω ≠ 0 := omega0_pos.ne' diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index 983753c1f9a2c..31c0ed062c498 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -779,9 +779,6 @@ theorem card_omega0 : card ω = ℵ₀ := theorem lift_omega0 : lift ω = ω := lift_lift _ -@[deprecated (since := "2024-09-30")] -alias lift_omega := lift_omega0 - /-! ### Definition and first properties of addition on ordinals From 5750078e996e08ab79ccc775c91e46ed17a803c3 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Mon, 18 Nov 2024 05:00:13 +0000 Subject: [PATCH 044/829] fix: linarith handling of decimals (#19174) [Zulip](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/linarith.20regression.20on.20real.20numbers) --- Mathlib/Analysis/Normed/Ring/Units.lean | 5 +---- .../Additive/AP/Three/Behrend.lean | 4 +--- Mathlib/Tactic/CancelDenoms/Core.lean | 18 ++++++++++++------ Mathlib/Tactic/Linarith/Preprocessing.lean | 3 ++- MathlibTest/linarith.lean | 10 ++++++++++ 5 files changed, 26 insertions(+), 14 deletions(-) diff --git a/Mathlib/Analysis/Normed/Ring/Units.lean b/Mathlib/Analysis/Normed/Ring/Units.lean index 5909da99e9e35..e0b887991a865 100644 --- a/Mathlib/Analysis/Normed/Ring/Units.lean +++ b/Mathlib/Analysis/Normed/Ring/Units.lean @@ -138,16 +138,13 @@ theorem inverse_one_sub_norm : (fun t : R => inverse (1 - t)) =O[𝓝 0] (fun _t simp only [IsBigO, IsBigOWith, Metric.eventually_nhds_iff] refine ⟨‖(1 : R)‖ + 1, (2 : ℝ)⁻¹, by norm_num, fun t ht ↦ ?_⟩ rw [dist_zero_right] at ht - have ht' : ‖t‖ < 1 := by - have : (2 : ℝ)⁻¹ < 1 := by cancel_denoms - linarith + have ht' : ‖t‖ < 1 := by linarith simp only [inverse_one_sub t ht', norm_one, mul_one, Set.mem_setOf_eq] change ‖∑' n : ℕ, t ^ n‖ ≤ _ have := tsum_geometric_le_of_norm_lt_one t ht' have : (1 - ‖t‖)⁻¹ ≤ 2 := by rw [← inv_inv (2 : ℝ)] refine inv_anti₀ (by norm_num) ?_ - have : (2 : ℝ)⁻¹ + (2 : ℝ)⁻¹ = 1 := by ring linarith linarith diff --git a/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean b/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean index d8b6d29880136..19ecd990580b6 100644 --- a/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean +++ b/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean @@ -282,8 +282,7 @@ theorem log_two_mul_two_le_sqrt_log_eight : log 2 * 2 ≤ √(log 8) := by theorem two_div_one_sub_two_div_e_le_eight : 2 / (1 - 2 / exp 1) ≤ 8 := by rw [div_le_iff₀, mul_sub, mul_one, mul_div_assoc', le_sub_comm, div_le_iff₀ (exp_pos _)] - · have : 16 < 6 * (2.7182818283 : ℝ) := by norm_num - linarith [exp_one_gt_d9] + · linarith [exp_one_gt_d9] rw [sub_pos, div_lt_one] <;> exact exp_one_gt_d9.trans' (by norm_num) theorem le_sqrt_log (hN : 4096 ≤ N) : log (2 / (1 - 2 / exp 1)) * (69 / 50) ≤ √(log ↑N) := by @@ -464,7 +463,6 @@ theorem exp_four_lt : exp 4 < 64 := by theorem four_zero_nine_six_lt_exp_sixteen : 4096 < exp 16 := by rw [← log_lt_iff_lt_exp (show (0 : ℝ) < 4096 by norm_num), show (4096 : ℝ) = 2 ^ 12 by norm_cast, ← rpow_natCast, log_rpow zero_lt_two, cast_ofNat] - have : 12 * (0.6931471808 : ℝ) < 16 := by norm_num linarith [log_two_lt_d9] theorem lower_bound_le_one' (hN : 2 ≤ N) (hN' : N ≤ 4096) : diff --git a/Mathlib/Tactic/CancelDenoms/Core.lean b/Mathlib/Tactic/CancelDenoms/Core.lean index 89310275c745f..b8ada93499ee9 100644 --- a/Mathlib/Tactic/CancelDenoms/Core.lean +++ b/Mathlib/Tactic/CancelDenoms/Core.lean @@ -219,6 +219,10 @@ def deriveThms : List Name := /-- Helper lemma to chain together a `simp` proof and the result of `mkProdPrf`. -/ theorem derive_trans {α} [Mul α] {a b c d : α} (h : a = b) (h' : c * b = d) : c * a = d := h ▸ h' +/-- Helper lemma to chain together two `simp` proofs and the result of `mkProdPrf`. -/ +theorem derive_trans₂ {α} [Mul α] {a b c d e : α} (h : a = b) (h' : b = c) (h'' : d * c = e) : + d * a = e := h ▸ h' ▸ h'' + /-- Given `e`, a term with rational division, produces a natural number `n` and a proof of `n*e = e'`, where `e'` has no division. Assumes "well-behaved" division. @@ -227,18 +231,20 @@ def derive (e : Expr) : MetaM (ℕ × Expr) := do trace[CancelDenoms] "e = {e}" let eSimp ← simpOnlyNames (config := Simp.neutralConfig) deriveThms e trace[CancelDenoms] "e simplified = {eSimp.expr}" - let (n, t) := findCancelFactor eSimp.expr - let ⟨u, tp, e⟩ ← inferTypeQ' eSimp.expr + let eSimpNormNum ← Mathlib.Meta.NormNum.deriveSimp {} false eSimp.expr + trace[CancelDenoms] "e norm_num'd = {eSimpNormNum.expr}" + let (n, t) := findCancelFactor eSimpNormNum.expr + let ⟨u, tp, e⟩ ← inferTypeQ' eSimpNormNum.expr let stp : Q(Field $tp) ← synthInstanceQ q(Field $tp) try have n' := (← mkOfNat tp q(inferInstance) <| mkRawNatLit <| n).1 let r ← mkProdPrf tp stp n n' t e trace[CancelDenoms] "pf : {← inferType r.pf}" let pf' ← - if let some pfSimp := eSimp.proof? then - mkAppM ``derive_trans #[pfSimp, r.pf] - else - pure r.pf + match eSimp.proof?, eSimpNormNum.proof? with + | some pfSimp, some pfSimp' => mkAppM ``derive_trans₂ #[pfSimp, pfSimp', r.pf] + | some pfSimp, none | none, some pfSimp => mkAppM ``derive_trans #[pfSimp, r.pf] + | none, none => pure r.pf return (n, pf') catch E => do throwError "CancelDenoms.derive failed to normalize {e}.\n{E.toMessageData}" diff --git a/Mathlib/Tactic/Linarith/Preprocessing.lean b/Mathlib/Tactic/Linarith/Preprocessing.lean index 418da06138f72..b97fd25e72ce5 100644 --- a/Mathlib/Tactic/Linarith/Preprocessing.lean +++ b/Mathlib/Tactic/Linarith/Preprocessing.lean @@ -250,7 +250,8 @@ def cancelDenoms : Preprocessor where name := "cancel denominators" transform := fun pf => (do let (_, lhs) ← parseCompAndExpr (← inferType pf) - guard <| lhs.containsConst (fun n => n = ``HDiv.hDiv || n = ``Div.div) + guard <| lhs.containsConst <| fun n => + n = ``HDiv.hDiv || n = ``Div.div || n = ``Inv.inv || n == ``OfScientific.ofScientific pure [← normalizeDenominatorsLHS pf lhs]) <|> return [pf] end cancelDenoms diff --git a/MathlibTest/linarith.lean b/MathlibTest/linarith.lean index b599128dcf2da..3f0afc36d4512 100644 --- a/MathlibTest/linarith.lean +++ b/MathlibTest/linarith.lean @@ -92,6 +92,16 @@ example (x : Rat) (h : 0 < x) : 0 < x/2/3 := by linarith example (x : Rat) (h : 0 < x) : 0 < x/(2/3) := by linarith +variable {K : Type*} [LinearOrderedField K] + +example (a : K) (ha : 10 / (8 + 2) ≤ a) : 1 ≤ a := by linarith + +example (a : K) (ha : 10 / 10 ^ 1 ≤ a) : 1 ≤ a := by linarith + +example (a : K) (ha : 10⁻¹ * 10 ≤ a) : 1 ≤ a := by linarith + +example (a : K) (ha : 1.0 ≤ a) : 1 ≤ a := by linarith + end cancel_denoms example (a b c : Rat) (h2 : b + 2 > 3 + b) : False := by From 6e2b040a2716a7e9265d7f837d12e8d2189a9880 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Mon, 18 Nov 2024 05:29:37 +0000 Subject: [PATCH 045/829] perf: hard-code typeclass inference in linarith (#18714) This PR modifies the `linarith` tactic to rely on custom copies of many lemmas, rather than the library versions of these lemmas. The difference is that the copies use high-level typeclasses (such as `OrderedSemiring`) rather than lower-level typeclasses (such as `AddRightMono`). This seems to improve `linarith` performance. The PR also introduces a performance test to the `linarith` test file, with a heartbeat limit which failed before this PR and passes (with leeway) after this PR. [Zulip](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/typeclass.20inference.20in.20linarith) --- Mathlib/Tactic/Linarith/Lemmas.lean | 24 ++++++++++++++++++++++ Mathlib/Tactic/Linarith/Preprocessing.lean | 4 ++-- Mathlib/Tactic/Linarith/Verification.lean | 10 ++++----- MathlibTest/linarith.lean | 20 ++++++++++++++++++ 4 files changed, 51 insertions(+), 7 deletions(-) diff --git a/Mathlib/Tactic/Linarith/Lemmas.lean b/Mathlib/Tactic/Linarith/Lemmas.lean index 452a98c48d2bb..5802a9ec5a4da 100644 --- a/Mathlib/Tactic/Linarith/Lemmas.lean +++ b/Mathlib/Tactic/Linarith/Lemmas.lean @@ -24,6 +24,8 @@ namespace Linarith universe u theorem lt_irrefl {α : Type u} [Preorder α] {a : α} : ¬a < a := _root_.lt_irrefl a +theorem zero_lt_one {α : Type u} [StrictOrderedSemiring α] : (0:α) < 1 := _root_.zero_lt_one + theorem eq_of_eq_of_eq {α} [OrderedSemiring α] {a b : α} (ha : a = 0) (hb : b = 0) : a + b = 0 := by simp [*] @@ -39,6 +41,22 @@ theorem le_of_le_of_eq {α} [OrderedSemiring α] {a b : α} (ha : a ≤ 0) (hb : theorem lt_of_lt_of_eq {α} [OrderedSemiring α] {a b : α} (ha : a < 0) (hb : b = 0) : a + b < 0 := by simp [*] +theorem add_nonpos {α : Type*} [OrderedSemiring α] {a b : α} (ha : a ≤ 0) (hb : b ≤ 0) : + a + b ≤ 0 := + _root_.add_nonpos ha hb + +theorem add_lt_of_le_of_neg {α : Type*} [StrictOrderedSemiring α] {a b c : α} (hbc : b ≤ c) + (ha : a < 0) : b + a < c := + _root_.add_lt_of_le_of_neg hbc ha + +theorem add_lt_of_neg_of_le {α : Type*} [StrictOrderedSemiring α] {a b c : α} (ha : a < 0) + (hbc : b ≤ c) : a + b < c := + _root_.add_lt_of_neg_of_le ha hbc + +theorem add_neg {α : Type*} [StrictOrderedSemiring α] {a b : α} (ha : a < 0) + (hb : b < 0) : a + b < 0 := + _root_.add_neg ha hb + theorem mul_neg {α} [StrictOrderedRing α] {a b : α} (ha : a < 0) (hb : 0 < b) : b * a < 0 := have : (-b)*a > 0 := mul_pos_of_neg_of_neg (neg_neg_of_pos hb) ha neg_of_neg_pos (by simpa) @@ -59,6 +77,12 @@ def _root_.Mathlib.Ineq.toConstMulName : Ineq → Lean.Name | .le => ``mul_nonpos | .eq => ``mul_eq +theorem sub_nonpos_of_le {α : Type*} [OrderedRing α] {a b : α} : a ≤ b → a - b ≤ 0 := + _root_.sub_nonpos_of_le + +theorem sub_neg_of_lt {α : Type*} [OrderedRing α] {a b : α} : a < b → a - b < 0 := + _root_.sub_neg_of_lt + lemma eq_of_not_lt_of_not_gt {α} [LinearOrder α] (a b : α) (h1 : ¬ a < b) (h2 : ¬ b < a) : a = b := le_antisymm (le_of_not_gt h2) (le_of_not_gt h1) diff --git a/Mathlib/Tactic/Linarith/Preprocessing.lean b/Mathlib/Tactic/Linarith/Preprocessing.lean index b97fd25e72ce5..eca9903c6a74d 100644 --- a/Mathlib/Tactic/Linarith/Preprocessing.lean +++ b/Mathlib/Tactic/Linarith/Preprocessing.lean @@ -205,8 +205,8 @@ and turns it into a proof of a comparison `_ R 0`, where `R ∈ {=, ≤, <}`. -/ partial def rearrangeComparison (e : Expr) : MetaM (Option Expr) := do match ← (← inferType e).ineq? with - | (Ineq.le, _) => try? <| mkAppM ``sub_nonpos_of_le #[e] - | (Ineq.lt, _) => try? <| mkAppM ``sub_neg_of_lt #[e] + | (Ineq.le, _) => try? <| mkAppM ``Linarith.sub_nonpos_of_le #[e] + | (Ineq.lt, _) => try? <| mkAppM ``Linarith.sub_neg_of_lt #[e] | (Ineq.eq, _) => try? <| mkAppM ``sub_eq_zero_of_eq #[e] /-- diff --git a/Mathlib/Tactic/Linarith/Verification.lean b/Mathlib/Tactic/Linarith/Verification.lean index 9cd251d4543f9..d9d9128bfee29 100644 --- a/Mathlib/Tactic/Linarith/Verification.lean +++ b/Mathlib/Tactic/Linarith/Verification.lean @@ -90,11 +90,11 @@ def addIneq : Ineq → Ineq → (Name × Ineq) | eq, le => (``Linarith.le_of_eq_of_le, le) | eq, lt => (``Linarith.lt_of_eq_of_lt, lt) | le, eq => (``Linarith.le_of_le_of_eq, le) - | le, le => (``add_nonpos, le) - | le, lt => (``add_lt_of_le_of_neg, lt) + | le, le => (``Linarith.add_nonpos, le) + | le, lt => (``Linarith.add_lt_of_le_of_neg, lt) | lt, eq => (``Linarith.lt_of_lt_of_eq, lt) - | lt, le => (``add_lt_of_neg_of_le, lt) - | lt, lt => (``Left.add_neg, lt) + | lt, le => (``Linarith.add_lt_of_neg_of_le, lt) + | lt, lt => (``Linarith.add_neg, lt) /-- `mkLTZeroProof coeffs pfs` takes a list of proofs of the form `tᵢ Rᵢ 0`, @@ -136,7 +136,7 @@ def typeOfIneqProof (prf : Expr) : MetaM Expr := do where the numerals are natively of type `tp`. -/ def mkNegOneLtZeroProof (tp : Expr) : MetaM Expr := do - let zero_lt_one ← mkAppOptM ``zero_lt_one #[tp, none, none, none, none, none] + let zero_lt_one ← mkAppOptM ``Linarith.zero_lt_one #[tp, none] mkAppM `neg_neg_of_pos #[zero_lt_one] /-- diff --git a/MathlibTest/linarith.lean b/MathlibTest/linarith.lean index 3f0afc36d4512..2341dfcd13359 100644 --- a/MathlibTest/linarith.lean +++ b/MathlibTest/linarith.lean @@ -8,6 +8,12 @@ private axiom test_sorry : ∀ {α}, α set_option linter.unusedVariables false set_option autoImplicit false +open Lean.Elab.Tactic in +def testSorryTac : TacticM Unit := do + let e ← getMainTarget + let t ← `(test_sorry) + closeMainGoalUsing `sorry fun _ _ => elabTerm t e + example {α} [LinearOrderedCommRing α] {a b : α} (h : a < b) (w : b < a) : False := by linarith @@ -275,6 +281,20 @@ example (x y z : ℚ) (hx : x < 5) (hx2 : x > 5) (hy : y < 5000000000) (hz : z > example (x y z : ℚ) (hx : x < 5) (hy : y < 5000000000) (hz : z > 34*y) : x ≤ 5 := by linarith only [hx] +/- The speed of `linarith` is very sensitive to how much typeclass inference is demanded by the +lemmas it orchestrates. + +This example took 3318 heartbeats (and 110 ms on a good laptop) on an earlier implementation, which +in several places used off-the-shelf library lemmas requiring low-level typeclasses, e.g. +`Add{Left,Right}{Strict}Mono`. + +After a tweak to rely only on custom `linarith` clones of these lemmas taking high-level typeclasses, +this now (November 2024) takes 1647 heartbeats (63 ms on a good laptop). -/ +set_option maxHeartbeats 2000 in +example {K : Type*} [LinearOrderedField K] {x y : K} + (h : x + y = 10) (h' : x + 2 * y ≤ 18) (h : x < 2) : False := by + linarith (config := {discharger := testSorryTac}) + example (u v x y A B : ℚ) (a : 0 < A) (a_1 : 0 <= 1 - A) From d2568da6e252a1bb378b585d39dcf12dd814f946 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Mon, 18 Nov 2024 06:07:26 +0000 Subject: [PATCH 046/829] feat(Topology/Sheaves/Stalks): Generalize to `stalkPushforward_iso_of_inducing` (#18268) port of https://github.com/leanprover-community/mathlib3/pull/17539 Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- .../Geometry/RingedSpace/OpenImmersion.lean | 4 +- .../Topology/Category/TopCat/OpenNhds.lean | 23 +++++++ Mathlib/Topology/Category/TopCat/Opens.lean | 61 +++++++++++++++++-- Mathlib/Topology/Sheaves/Stalks.lean | 36 +++-------- 4 files changed, 92 insertions(+), 32 deletions(-) diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index cd6872a8c7543..b937022700288 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -737,8 +737,8 @@ theorem of_stalk_iso {X Y : SheafedSpace C} (f : X ⟶ Y) (hf : IsOpenEmbedding rintro ⟨_, y, hy, rfl⟩ specialize H y delta PresheafedSpace.Hom.stalkMap at H - haveI H' := - TopCat.Presheaf.stalkPushforward.stalkPushforward_iso_of_isOpenEmbedding C hf X.presheaf y + haveI H' := TopCat.Presheaf.stalkPushforward.stalkPushforward_iso_of_isInducing C + hf.toIsInducing X.presheaf y have := IsIso.comp_isIso' H (@IsIso.inv_isIso _ _ _ _ _ H') rwa [Category.assoc, IsIso.hom_inv_id, Category.comp_id] at this } diff --git a/Mathlib/Topology/Category/TopCat/OpenNhds.lean b/Mathlib/Topology/Category/TopCat/OpenNhds.lean index d18408ac7ca35..982cb295c7d5b 100644 --- a/Mathlib/Topology/Category/TopCat/OpenNhds.lean +++ b/Mathlib/Topology/Category/TopCat/OpenNhds.lean @@ -152,3 +152,26 @@ def adjunctionNhds (h : IsOpenMap f) (x : X) : IsOpenMap.functorNhds h x ⊣ Ope counit := { app := fun _ => homOfLE fun _ ⟨_, hfxV, hxy⟩ => hxy ▸ hfxV } end IsOpenMap + +namespace Topology.IsInducing + +open TopologicalSpace + +variable {f} + +/-- An inducing map `f : X ⟶ Y` induces a functor `open_nhds x ⥤ open_nhds (f x)`. -/ +@[simps] +def functorNhds (h : IsInducing f) (x : X) : + OpenNhds x ⥤ OpenNhds (f x) where + obj U := ⟨h.functor.obj U.1, (h.mem_functorObj_iff U.1).mpr U.2⟩ + map := h.functor.map + +/-- +An inducing map `f : X ⟶ Y` induces an adjunction between `open_nhds x` and `open_nhds (f x)`. +-/ +def adjunctionNhds (h : IsInducing f) (x : X) : + OpenNhds.map f x ⊣ h.functorNhds x where + unit := { app := fun U => homOfLE (h.adjunction.unit.app U.1).le } + counit := { app := fun U => homOfLE (h.adjunction.counit.app U.1).le } + +end Topology.IsInducing diff --git a/Mathlib/Topology/Category/TopCat/Opens.lean b/Mathlib/Topology/Category/TopCat/Opens.lean index 20109781231aa..a34c444596d5a 100644 --- a/Mathlib/Topology/Category/TopCat/Opens.lean +++ b/Mathlib/Topology/Category/TopCat/Opens.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ -import Mathlib.CategoryTheory.Category.Preorder +import Mathlib.CategoryTheory.Category.GaloisConnection import Mathlib.CategoryTheory.EqToHom import Mathlib.Topology.Category.TopCat.EpiMono import Mathlib.Topology.Sets.Opens @@ -21,7 +21,7 @@ but using it results in unresolvable `Eq.rec` terms in goals.) Really it's a 2-functor from (spaces, continuous functions, equalities) to (categories, functors, natural isomorphisms). We don't attempt to set up the full theory here, but do provide the natural isomorphisms -`mapId : map (𝟙 X) ≅ 𝟭 (opens X)` and +`mapId : map (𝟙 X) ≅ 𝟭 (Opens X)` and `mapComp : map (f ≫ g) ≅ map g ⋙ map f`. Beyond that, there's a collection of simp lemmas for working with these constructions. @@ -94,7 +94,7 @@ theorem leSupr_apply_mk {ι : Type*} (U : ι → Opens X) (i : ι) (x) (m) : (leSupr U i) ⟨x, m⟩ = ⟨x, (le_iSup U i : _) m⟩ := rfl -/-- The functor from open sets in `X` to `Top`, +/-- The functor from open sets in `X` to `TopCat`, realising each open set as a topological space itself. -/ def toTopCat (X : TopCat.{u}) : Opens X ⥤ TopCat where @@ -281,7 +281,7 @@ def IsOpenMap.functor {X Y : TopCat} {f : X ⟶ Y} (hf : IsOpenMap f) : Opens X /-- An open map `f : X ⟶ Y` induces an adjunction between `Opens X` and `Opens Y`. -/ def IsOpenMap.adjunction {X Y : TopCat} {f : X ⟶ Y} (hf : IsOpenMap f) : - Adjunction hf.functor (TopologicalSpace.Opens.map f) where + hf.functor ⊣ Opens.map f where unit := { app := fun _ => homOfLE fun x hxU => ⟨x, hxU, rfl⟩ } counit := { app := fun _ => homOfLE fun _ ⟨_, hfxV, hxy⟩ => hxy ▸ hfxV } @@ -302,6 +302,59 @@ lemma Topology.IsOpenEmbedding.functor_obj_injective {X Y : TopCat} {f : X ⟶ Y @[deprecated (since := "2024-10-18")] alias OpenEmbedding.functor_obj_injective := IsOpenEmbedding.functor_obj_injective +namespace Topology.IsInducing + +/-- Given an inducing map `X ⟶ Y` and some `U : Opens X`, this is the union of all open sets +whose preimage is `U`. This is right adjoint to `Opens.map`. -/ +@[nolint unusedArguments] +def functorObj {X Y : TopCat} {f : X ⟶ Y} (_ : IsInducing f) (U : Opens X) : Opens Y := + sSup { s : Opens Y | (Opens.map f).obj s = U } + +lemma map_functorObj {X Y : TopCat} {f : X ⟶ Y} (hf : IsInducing f) + (U : Opens X) : + (Opens.map f).obj (hf.functorObj U) = U := by + apply le_antisymm + · rintro x ⟨_, ⟨s, rfl⟩, _, ⟨rfl : _ = U, rfl⟩, hx : f x ∈ s⟩; exact hx + · intros x hx + obtain ⟨U, hU⟩ := U + obtain ⟨t, ht, rfl⟩ := hf.isOpen_iff.mp hU + exact Opens.mem_sSup.mpr ⟨⟨_, ht⟩, rfl, hx⟩ + +lemma mem_functorObj_iff {X Y : TopCat} {f : X ⟶ Y} (hf : IsInducing f) (U : Opens X) + {x : X} : f x ∈ hf.functorObj U ↔ x ∈ U := by + conv_rhs => rw [← hf.map_functorObj U] + rfl + +lemma le_functorObj_iff {X Y : TopCat} {f : X ⟶ Y} (hf : IsInducing f) {U : Opens X} + {V : Opens Y} : V ≤ hf.functorObj U ↔ (Opens.map f).obj V ≤ U := by + obtain ⟨U, hU⟩ := U + obtain ⟨t, ht, rfl⟩ := hf.isOpen_iff.mp hU + constructor + · exact fun i x hx ↦ (hf.mem_functorObj_iff ((Opens.map f).obj ⟨t, ht⟩)).mp (i hx) + · intros h x hx + refine Opens.mem_sSup.mpr ⟨⟨_, V.2.union ht⟩, Opens.ext ?_, Set.mem_union_left t hx⟩ + dsimp + rwa [Set.union_eq_right] + +/-- An inducing map `f : X ⟶ Y` induces a Galois insertion between `Opens Y` and `Opens X`. -/ +def opensGI {X Y : TopCat} {f : X ⟶ Y} (hf : IsInducing f) : + GaloisInsertion (Opens.map f).obj hf.functorObj := + ⟨_, fun _ _ ↦ hf.le_functorObj_iff.symm, fun U ↦ (hf.map_functorObj U).ge, fun _ _ ↦ rfl⟩ + +/-- An inducing map `f : X ⟶ Y` induces a functor `Opens X ⥤ Opens Y`. -/ +@[simps] +def functor {X Y : TopCat} {f : X ⟶ Y} (hf : IsInducing f) : + Opens X ⥤ Opens Y where + obj := hf.functorObj + map {U V} h := homOfLE (hf.le_functorObj_iff.mpr ((hf.map_functorObj U).trans_le h.le)) + +/-- An inducing map `f : X ⟶ Y` induces an adjunction between `Opens Y` and `Opens X`. -/ +def adjunction {X Y : TopCat} {f : X ⟶ Y} (hf : IsInducing f) : + Opens.map f ⊣ hf.functor := + hf.opensGI.gc.adjunction + +end Topology.IsInducing + namespace TopologicalSpace.Opens open TopologicalSpace diff --git a/Mathlib/Topology/Sheaves/Stalks.lean b/Mathlib/Topology/Sheaves/Stalks.lean index 2aebc0ddd16a7..8c4b279f06b5d 100644 --- a/Mathlib/Topology/Sheaves/Stalks.lean +++ b/Mathlib/Topology/Sheaves/Stalks.lean @@ -204,34 +204,18 @@ theorem comp (ℱ : X.Presheaf C) (f : X ⟶ Y) (g : Y ⟶ Z) (x : X) : ext simp [germ, stalkPushforward] -theorem stalkPushforward_iso_of_isOpenEmbedding {f : X ⟶ Y} (hf : IsOpenEmbedding f) +theorem stalkPushforward_iso_of_isInducing {f : X ⟶ Y} (hf : IsInducing f) (F : X.Presheaf C) (x : X) : IsIso (F.stalkPushforward _ f x) := by - haveI := Functor.initial_of_adjunction (hf.isOpenMap.adjunctionNhds x) - convert - ((Functor.Final.colimitIso (hf.isOpenMap.functorNhds x).op - ((OpenNhds.inclusion (f x)).op ⋙ f _* F) : - _).symm ≪≫ - colim.mapIso _).isIso_hom - swap - · fapply NatIso.ofComponents - · intro U - refine F.mapIso (eqToIso ?_) - dsimp only [Functor.op] - exact congr_arg op (Opens.ext <| Set.preimage_image_eq (unop U).1.1 hf.injective) - · intro U V i; erw [← F.map_comp, ← F.map_comp]; congr 1 - · change (_ : colimit _ ⟶ _) = (_ : colimit _ ⟶ _) - ext U - rw [← Iso.comp_inv_eq] - erw [colimit.ι_map_assoc] - rw [colimit.ι_pre, Category.assoc] - erw [colimit.ι_map_assoc, colimit.ι_pre, ← F.map_comp_assoc] - apply colimit.w ((OpenNhds.inclusion (f x)).op ⋙ f _* F) _ - dsimp only [Functor.op] - refine ((homOfLE ?_).op : op (unop U) ⟶ _) - exact Set.image_preimage_subset _ _ - + haveI := Functor.initial_of_adjunction (hf.adjunctionNhds x) + convert (Functor.Final.colimitIso (OpenNhds.map f x).op ((OpenNhds.inclusion x).op ⋙ F)).isIso_hom + refine stalk_hom_ext _ fun U hU ↦ (stalkPushforward_germ _ f F _ x hU).trans ?_ + symm + exact colimit.ι_pre ((OpenNhds.inclusion x).op ⋙ F) (OpenNhds.map f x).op _ + +@[deprecated (since := "2024-10-27")] +alias stalkPushforward_iso_of_isOpenEmbedding := stalkPushforward_iso_of_isInducing @[deprecated (since := "2024-10-18")] -alias stalkPushforward_iso_of_openEmbedding := stalkPushforward_iso_of_isOpenEmbedding +alias stalkPushforward_iso_of_openEmbedding := stalkPushforward_iso_of_isInducing end stalkPushforward From 0002a67b034519255e84fd3061d9c93c0bfcc5a4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 18 Nov 2024 06:50:52 +0000 Subject: [PATCH 047/829] =?UTF-8?q?feat(SetTheory/Ordinal/Basic):=20`?= =?UTF-8?q?=CF=89=20=E2=89=A4=20a.ord=20=E2=86=94=20=E2=84=B5=E2=82=80=20?= =?UTF-8?q?=E2=89=A4=20a`=20(#18901)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yury G. Kudryashov --- Mathlib/SetTheory/Ordinal/Basic.lean | 28 ++++++++++++++++++++++++++++ 1 file changed, 28 insertions(+) diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index 31c0ed062c498..65e5733c47823 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -1232,6 +1232,26 @@ theorem ord_eq_zero {a : Cardinal} : a.ord = 0 ↔ a = 0 := theorem ord_eq_one {a : Cardinal} : a.ord = 1 ↔ a = 1 := ord_injective.eq_iff' ord_one +@[simp] +theorem omega0_le_ord {a : Cardinal} : ω ≤ a.ord ↔ ℵ₀ ≤ a := by + rw [← ord_aleph0, ord_le_ord] + +@[simp] +theorem ord_le_omega0 {a : Cardinal} : a.ord ≤ ω ↔ a ≤ ℵ₀ := by + rw [← ord_aleph0, ord_le_ord] + +@[simp] +theorem ord_lt_omega0 {a : Cardinal} : a.ord < ω ↔ a < ℵ₀ := + le_iff_le_iff_lt_iff_lt.1 omega0_le_ord + +@[simp] +theorem omega0_lt_ord {a : Cardinal} : ω < a.ord ↔ ℵ₀ < a := + le_iff_le_iff_lt_iff_lt.1 ord_le_omega0 + +@[simp] +theorem ord_eq_omega0 {a : Cardinal} : a.ord = ω ↔ a = ℵ₀ := + ord_injective.eq_iff' ord_aleph0 + /-- The ordinal corresponding to a cardinal `c` is the least ordinal whose cardinal is `c`. This is the order-embedding version. For the regular function, see `ord`. -/ @@ -1325,6 +1345,14 @@ theorem ofNat_le_card {o} {n : ℕ} [n.AtLeastTwo] : (no_index (OfNat.ofNat n : Cardinal)) ≤ card o ↔ (OfNat.ofNat n : Ordinal) ≤ o := nat_le_card +@[simp] +theorem aleph0_le_card {o} : ℵ₀ ≤ card o ↔ ω ≤ o := by + rw [← ord_le, ord_aleph0] + +@[simp] +theorem card_lt_aleph0 {o} : card o < ℵ₀ ↔ o < ω := + le_iff_le_iff_lt_iff_lt.1 aleph0_le_card + @[simp] theorem nat_lt_card {o} {n : ℕ} : (n : Cardinal) < card o ↔ (n : Ordinal) < o := by rw [← succ_le_iff, ← succ_le_iff, ← nat_succ, nat_le_card] From a2d097b70623b372264bdc073d89bc1164205341 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 18 Nov 2024 07:12:19 +0000 Subject: [PATCH 048/829] chore: more uses of `calc` and `gcongr` (#19173) Motivated by the deprecation in #19081 --- .../Order/GroupWithZero/Unbundled.lean | 3 + .../Fourier/FourierTransformDeriv.lean | 4 +- Mathlib/Analysis/InnerProductSpace/Basic.lean | 2 +- Mathlib/Analysis/Normed/Field/Ultra.lean | 3 +- Mathlib/Analysis/SpecialFunctions/Exp.lean | 2 +- .../SpecialFunctions/JapaneseBracket.lean | 17 +++--- .../SimpleGraph/Regularity/Bound.lean | 2 +- .../SimpleGraph/Regularity/Chunk.lean | 56 +++++++++---------- .../SimpleGraph/Regularity/Lemma.lean | 5 +- Mathlib/Data/Complex/Exponential.lean | 13 ++--- Mathlib/Data/Real/Archimedean.lean | 15 ++--- Mathlib/Data/Real/Basic.lean | 13 +++-- Mathlib/Data/Real/Pi/Bounds.lean | 29 ++++------ .../EisensteinSeries/UniformConvergence.lean | 6 +- .../ModularForms/JacobiTheta/Bounds.lean | 4 +- .../NumberField/Units/DirichletTheorem.lean | 5 +- 16 files changed, 83 insertions(+), 96 deletions(-) diff --git a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean index 8fa3d0864cc4a..f11dc3c1af754 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean @@ -1033,6 +1033,9 @@ theorem pow_le_pow_left₀ (ha : 0 ≤ a) (hab : a ≤ b) : ∀ n, a ^ n ≤ b ^ | n + 1 => by simpa only [pow_succ'] using mul_le_mul hab (pow_le_pow_left₀ ha hab _) (pow_nonneg ha _) (ha.trans hab) +lemma pow_left_monotoneOn : MonotoneOn (fun a : M₀ ↦ a ^ n) {x | 0 ≤ x} := + fun _a ha _b _ hab ↦ pow_le_pow_left₀ ha hab _ + end variable [Preorder α] {f g : α → M₀} diff --git a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean index 3412809b2fb10..7a53f32811c4c 100644 --- a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean +++ b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean @@ -412,8 +412,8 @@ lemma norm_iteratedFDeriv_fourierPowSMulRight gcongr with i hi · rw [← Nat.cast_pow, Nat.cast_le] calc n.descFactorial i ≤ n ^ i := Nat.descFactorial_le_pow _ _ - _ ≤ (n + 1) ^ i := pow_le_pow_left (by omega) (by omega) i - _ ≤ (n + 1) ^ k := pow_right_mono₀ (by omega) (Finset.mem_range_succ_iff.mp hi) + _ ≤ (n + 1) ^ i := by gcongr; omega + _ ≤ (n + 1) ^ k := by gcongr; exacts [le_add_self, Finset.mem_range_succ_iff.mp hi] · exact hv _ (by omega) _ (by omega) _ = (2 * n + 2) ^ k * (‖L‖^n * C) := by simp only [← Finset.sum_mul, ← Nat.cast_sum, Nat.sum_range_choose, mul_one, ← mul_assoc, diff --git a/Mathlib/Analysis/InnerProductSpace/Basic.lean b/Mathlib/Analysis/InnerProductSpace/Basic.lean index 86baea0a46b96..4acfe80398b10 100644 --- a/Mathlib/Analysis/InnerProductSpace/Basic.lean +++ b/Mathlib/Analysis/InnerProductSpace/Basic.lean @@ -1212,7 +1212,7 @@ instance (priority := 100) InnerProductSpace.toUniformConvexSpace : UniformConve refine le_sqrt_of_sq_le ?_ rw [sq, eq_sub_iff_add_eq.2 (parallelogram_law_with_norm ℝ x y), ← sq ‖x - y‖, hx, hy] ring_nf - exact sub_le_sub_left (pow_le_pow_left hε.le hxy _) 4⟩ + gcongr⟩ section Complex_Seminormed diff --git a/Mathlib/Analysis/Normed/Field/Ultra.lean b/Mathlib/Analysis/Normed/Field/Ultra.lean index 0f85eefafcd44..8ef392b9b1afd 100644 --- a/Mathlib/Analysis/Normed/Field/Ultra.lean +++ b/Mathlib/Analysis/Normed/Field/Ultra.lean @@ -89,8 +89,7 @@ lemma isUltrametricDist_of_forall_pow_norm_le_nsmul_pow_max_one_norm -- `‖x + 1‖ ^ m ≤ (m + 1) • max 1 ‖x‖ ^ m`, so we're done -- we can distribute powers into the right term of `max` have hp : max 1 ‖x‖ ^ m = max 1 (‖x‖ ^ m) := by - have : MonotoneOn (fun a : ℝ ↦ a ^ m) (Set.Ici _) := fun a h b _ h' ↦ pow_le_pow_left h h' _ - rw [this.map_max (by simp [zero_le_one]) (norm_nonneg x), one_pow] + rw [pow_left_monotoneOn.map_max (by simp [zero_le_one]) (norm_nonneg x), one_pow] rw [hp] at hm refine le_of_pow_le_pow_left₀ (fun h ↦ ?_) (zero_lt_one.trans ha').le ((h _ _).trans hm.le) simp only [h, zero_add, pow_zero, max_self, one_smul, lt_self_iff_false] at hm diff --git a/Mathlib/Analysis/SpecialFunctions/Exp.lean b/Mathlib/Analysis/SpecialFunctions/Exp.lean index 2ce030d8b645e..543a930581f63 100644 --- a/Mathlib/Analysis/SpecialFunctions/Exp.lean +++ b/Mathlib/Analysis/SpecialFunctions/Exp.lean @@ -264,7 +264,7 @@ theorem tendsto_exp_div_pow_atTop (n : ℕ) : Tendsto (fun x => exp x / x ^ n) a have hx₀ : 0 < x := (Nat.cast_nonneg N).trans_lt hx rw [Set.mem_Ici, le_div_iff₀ (pow_pos hx₀ _), ← le_div_iff₀' hC₀] calc - x ^ n ≤ ⌈x⌉₊ ^ n := mod_cast pow_le_pow_left hx₀.le (Nat.le_ceil _) _ + x ^ n ≤ ⌈x⌉₊ ^ n := by gcongr; exact Nat.le_ceil _ _ ≤ exp ⌈x⌉₊ / (exp 1 * C) := mod_cast (hN _ (Nat.lt_ceil.2 hx).le).le _ ≤ exp (x + 1) / (exp 1 * C) := by gcongr; exact (Nat.ceil_lt_add_one hx₀.le).le _ = exp x / C := by rw [add_comm, exp_add, mul_div_mul_left _ _ (exp_pos _).ne'] diff --git a/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean b/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean index 52481419b2ca5..a785ffd1f169b 100644 --- a/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean +++ b/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean @@ -72,15 +72,14 @@ variable {E} theorem finite_integral_rpow_sub_one_pow_aux {r : ℝ} (n : ℕ) (hnr : (n : ℝ) < r) : (∫⁻ x : ℝ in Ioc 0 1, ENNReal.ofReal ((x ^ (-r⁻¹) - 1) ^ n)) < ∞ := by have hr : 0 < r := lt_of_le_of_lt n.cast_nonneg hnr - have h_int : ∀ x : ℝ, x ∈ Ioc (0 : ℝ) 1 → - ENNReal.ofReal ((x ^ (-r⁻¹) - 1) ^ n) ≤ ENNReal.ofReal (x ^ (-(r⁻¹ * n))) := fun x hx ↦ by - apply ENNReal.ofReal_le_ofReal - rw [← neg_mul, rpow_mul hx.1.le, rpow_natCast] - refine pow_le_pow_left ?_ (by simp only [sub_le_self_iff, zero_le_one]) n - rw [le_sub_iff_add_le', add_zero] - refine Real.one_le_rpow_of_pos_of_le_one_of_nonpos hx.1 hx.2 ?_ - rw [Right.neg_nonpos_iff, inv_nonneg] - exact hr.le + have h_int x (hx : x ∈ Ioc (0 : ℝ) 1) := by + calc + ENNReal.ofReal ((x ^ (-r⁻¹) - 1) ^ n) ≤ .ofReal ((x ^ (-r⁻¹) - 0) ^ n) := by + gcongr + · rw [sub_nonneg] + exact Real.one_le_rpow_of_pos_of_le_one_of_nonpos hx.1 hx.2 (by simpa using hr.le) + · norm_num + _ = .ofReal (x ^ (-(r⁻¹ * n))) := by simp [rpow_mul hx.1.le, ← neg_mul] refine lt_of_le_of_lt (setLIntegral_mono' measurableSet_Ioc h_int) ?_ refine IntegrableOn.setLIntegral_lt_top ?_ rw [← intervalIntegrable_iff_integrableOn_Ioc_of_le zero_le_one] diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean index 10296b92189c4..038a7f498f743 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean @@ -69,7 +69,7 @@ private theorem eps_pos {ε : ℝ} {n : ℕ} (h : 100 ≤ (4 : ℝ) ^ n * ε ^ 5 (pos_of_mul_pos_right ((show 0 < (100 : ℝ) by norm_num).trans_le h) (by positivity)) private theorem m_pos [Nonempty α] (hPα : #P.parts * 16 ^ #P.parts ≤ card α) : 0 < m := - Nat.div_pos ((Nat.mul_le_mul_left _ <| Nat.pow_le_pow_left (by norm_num) _).trans hPα) <| + Nat.div_pos (hPα.trans' <| by unfold stepBound; gcongr; norm_num) <| stepBound_pos (P.parts_nonempty <| univ_nonempty.ne_empty).card_pos /-- Local extension for the `positivity` tactic: A few facts that are needed many times for the diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean index 2f4239ee94ad3..7759eae0b284d 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean @@ -226,20 +226,19 @@ private theorem one_sub_le_m_div_m_add_one_sq [Nonempty α] sz_positivity private theorem m_add_one_div_m_le_one_add [Nonempty α] - (hPα : #P.parts * 16 ^ #P.parts ≤ card α) (hPε : ↑100 ≤ ↑4 ^ #P.parts * ε ^ 5) - (hε₁ : ε ≤ 1) : ((m + 1 : ℝ) / m) ^ 2 ≤ ↑1 + ε ^ 5 / 49 := by - rw [same_add_div] - swap; · sz_positivity - have : ↑1 + ↑1 / (m : ℝ) ≤ ↑1 + ε ^ 5 / 100 := by - rw [add_le_add_iff_left, ← one_div_div (100 : ℝ)] - exact one_div_le_one_div_of_le (by sz_positivity) (hundred_div_ε_pow_five_le_m hPα hPε) - refine (pow_le_pow_left ?_ this 2).trans ?_ - · positivity - rw [add_sq, one_pow, add_assoc, add_le_add_iff_left, mul_one, ← le_sub_iff_add_le', - 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₀ (by sz_positivity) hε₁).trans (by norm_num) + (hPα : #P.parts * 16 ^ #P.parts ≤ card α) (hPε : ↑100 ≤ ↑4 ^ #P.parts * ε ^ 5) (hε₁ : ε ≤ 1) : + ((m + 1 : ℝ) / m) ^ 2 ≤ ↑1 + ε ^ 5 / 49 := by + have : 0 ≤ ε := by sz_positivity + rw [same_add_div (by sz_positivity)] + calc + _ ≤ (1 + ε ^ 5 / 100) ^ 2 := by + gcongr (1 + ?_) ^ 2 + rw [← one_div_div (100 : ℝ)] + exact one_div_le_one_div_of_le (by sz_positivity) (hundred_div_ε_pow_five_le_m hPα hPε) + _ = 1 + ε ^ 5 * (50⁻¹ + ε ^ 5 / 10000) := by ring + _ ≤ 1 + ε ^ 5 * (50⁻¹ + 1 ^ 5 / 10000) := by gcongr + _ ≤ 1 + ε ^ 5 * 49⁻¹ := by gcongr; norm_num + _ = 1 + ε ^ 5 / 49 := by rw [div_eq_mul_inv] private theorem density_sub_eps_le_sum_density_div_card [Nonempty α] (hPα : #P.parts * 16 ^ #P.parts ≤ card α) (hPε : ↑100 ≤ ↑4 ^ #P.parts * ε ^ 5) @@ -346,21 +345,20 @@ private theorem edgeDensity_chunk_aux [Nonempty α] · exact mod_cast G.edgeDensity_le_one _ _ · exact div_le_div_of_nonneg_left (by sz_positivity) (by norm_num) (by norm_num) rw [← sub_nonneg] at hGε - have : ↑(G.edgeDensity U V) - ε ^ 5 / ↑50 ≤ - (∑ ab ∈ (chunk hP G ε hU).parts.product (chunk hP G ε hV).parts, - (G.edgeDensity ab.1 ab.2 : ℝ)) / ↑16 ^ #P.parts := by - have rflU := Set.Subset.refl (chunk hP G ε hU).parts.toSet - have rflV := Set.Subset.refl (chunk hP G ε hV).parts.toSet - refine (le_trans ?_ <| density_sub_eps_le_sum_density_div_card hPα hPε rflU rflV).trans ?_ - · rw [biUnion_parts, biUnion_parts] - · rw [card_chunk (m_pos hPα).ne', card_chunk (m_pos hPα).ne', ← cast_mul, ← mul_pow, cast_pow] - norm_cast - refine le_trans ?_ (pow_le_pow_left hGε this 2) - rw [sub_sq, sub_add, sub_le_sub_iff_left] - refine (sub_le_self _ <| sq_nonneg <| ε ^ 5 / 50).trans ?_ - rw [mul_right_comm, mul_div_left_comm, div_eq_mul_inv (ε ^ 5), - show (2 : ℝ) / 50 = 25⁻¹ by norm_num] - exact mul_le_of_le_one_right (by sz_positivity) (mod_cast G.edgeDensity_le_one _ _) + have : 0 ≤ ε := by sz_positivity + calc + _ = G.edgeDensity U V ^ 2 - 1 * ε ^ 5 / 25 + 0 ^ 10 / 2500 := by ring + _ ≤ G.edgeDensity U V ^ 2 - G.edgeDensity U V * ε ^ 5 / 25 + ε ^ 10 / 2500 := by + gcongr; exact mod_cast G.edgeDensity_le_one .. + _ = (G.edgeDensity U V - ε ^ 5 / 50) ^ 2 := by ring + _ ≤ _ := by + gcongr + have rflU := Set.Subset.refl (chunk hP G ε hU).parts.toSet + have rflV := Set.Subset.refl (chunk hP G ε hV).parts.toSet + refine (le_trans ?_ <| density_sub_eps_le_sum_density_div_card hPα hPε rflU rflV).trans ?_ + · rw [biUnion_parts, biUnion_parts] + · rw [card_chunk (m_pos hPα).ne', card_chunk (m_pos hPα).ne', ← cast_mul, ← mul_pow, cast_pow] + norm_cast private theorem abs_density_star_sub_density_le_eps (hPε : ↑100 ≤ ↑4 ^ #P.parts * ε ^ 5) (hε₁ : ε ≤ 1) {hU : U ∈ P.parts} {hV : V ∈ P.parts} (hUV' : U ≠ V) (hUV : ¬G.IsUniform ε U V) : diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Lemma.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Lemma.lean index dd9d05a8d0c75..bc460c101fb0e 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Lemma.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Lemma.lean @@ -100,8 +100,9 @@ theorem szemeredi_regularity (hε : 0 < ε) (hl : l ≤ card α) : obtain ⟨P, hP₁, hP₂, hP₃, hP₄⟩ := h (⌊4 / ε ^ 5⌋₊ + 1) refine ⟨P, hP₁, (le_initialBound _ _).trans hP₂, hP₃.trans ?_, hP₄.resolve_right fun hPenergy => lt_irrefl (1 : ℝ) ?_⟩ - · rw [iterate_succ_apply'] - exact mul_le_mul_left' (pow_le_pow_left (by norm_num) (by norm_num) _) _ + · rw [iterate_succ_apply', stepBound, bound] + gcongr + norm_num calc (1 : ℝ) = ε ^ 5 / ↑4 * (↑4 / ε ^ 5) := by rw [mul_comm, div_mul_div_cancel₀ (pow_pos hε 5).ne']; norm_num diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index 8cc3b019b5445..87be1b9c712d2 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -1441,13 +1441,12 @@ theorem one_sub_div_pow_le_exp_neg {n : ℕ} {t : ℝ} (ht' : t ≤ n) : (1 - t rcases eq_or_ne n 0 with (rfl | hn) · simp rwa [Nat.cast_zero] at ht' - convert pow_le_pow_left ?_ (one_sub_le_exp_neg (t / n)) n using 2 - · rw [← Real.exp_nat_mul] - congr 1 - field_simp - ring_nf - · rwa [sub_nonneg, div_le_one] - positivity + calc + (1 - t / n) ^ n ≤ rexp (-(t / n)) ^ n := by + gcongr + · exact sub_nonneg.2 <| div_le_one_of_le₀ ht' n.cast_nonneg + · exact one_sub_le_exp_neg _ + _ = rexp (-t) := by rw [← Real.exp_nat_mul, mul_neg, mul_comm, div_mul_cancel₀]; positivity end Real diff --git a/Mathlib/Data/Real/Archimedean.lean b/Mathlib/Data/Real/Archimedean.lean index 76c63d50d2371..55226e28d84b8 100644 --- a/Mathlib/Data/Real/Archimedean.lean +++ b/Mathlib/Data/Real/Archimedean.lean @@ -374,14 +374,11 @@ lemma exists_natCast_add_one_lt_pow_of_one_lt (ha : 1 < a) : ∃ m : ℕ, (m + 1 rw [le_tsub_iff_left hq.le] exact hq use 2 * k ^ 2 - refine (pow_lt_pow_left₀ hk (by positivity) (by simp [posk.ne'])).trans_le' ?_ - rcases k.zero_le.eq_or_lt with rfl|kpos - · simp - rw [pow_two, mul_left_comm, pow_mul] - have := mul_add_one_le_add_one_pow (a := 1 / k) (by simp) k - rw [div_mul_cancel₀ _ (by simp [kpos.ne'])] at this - refine (pow_le_pow_left (by positivity) this _).trans' ?_ - rw [mul_left_comm, ← pow_two] - exact_mod_cast Nat.two_mul_sq_add_one_le_two_pow_two_mul _ + calc + ((2 * k ^ 2 : ℕ) + 1 : ℝ) ≤ 2 ^ (2 * k) := mod_cast Nat.two_mul_sq_add_one_le_two_pow_two_mul _ + _ = (1 / k * k + 1 : ℝ) ^ (2 * k) := by simp [posk.ne']; norm_num + _ ≤ ((1 / k + 1) ^ k : ℝ) ^ (2 * k) := by gcongr; exact mul_add_one_le_add_one_pow (by simp) _ + _ = (1 / k + 1 : ℝ) ^ (2 * k ^ 2) := by rw [← pow_mul, mul_left_comm, sq] + _ < a ^ (2 * k ^ 2) := by gcongr end Real diff --git a/Mathlib/Data/Real/Basic.lean b/Mathlib/Data/Real/Basic.lean index 384f201b82a86..a47daa561e74e 100644 --- a/Mathlib/Data/Real/Basic.lean +++ b/Mathlib/Data/Real/Basic.lean @@ -584,11 +584,14 @@ lemma mul_add_one_le_add_one_pow {a : ℝ} (ha : 0 ≤ a) (b : ℕ) : a * b + 1 induction b generalizing a with | zero => simp | succ b hb => - rw [Nat.cast_add_one, mul_add, mul_one, add_right_comm, pow_succ, mul_add, mul_one, add_comm] - gcongr - · rw [le_mul_iff_one_le_left ha'] - exact (pow_le_pow_left zero_le_one (by simpa using ha'.le) b).trans' (by simp) - · exact hb ha' + calc + a * ↑(b + 1) + 1 = (0 + 1) ^ b * a + (a * b + 1) := by + simp [mul_add, add_assoc, add_left_comm] + _ ≤ (a + 1) ^ b * a + (a + 1) ^ b := by + gcongr + · norm_num + · exact hb ha' + _ = (a + 1) ^ (b + 1) := by simp [pow_succ, mul_add] end Real diff --git a/Mathlib/Data/Real/Pi/Bounds.lean b/Mathlib/Data/Real/Pi/Bounds.lean index 131af1b4136f2..d04c703754a53 100644 --- a/Mathlib/Data/Real/Pi/Bounds.lean +++ b/Mathlib/Data/Real/Pi/Bounds.lean @@ -34,25 +34,16 @@ 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] - 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₀'] - · refine le_trans pi_le_four ?_ - simp only [show (4 : ℝ) = (2 : ℝ) ^ 2 by norm_num, mul_one] - apply pow_right_mono₀ (by norm_num) - apply le_add_of_nonneg_left; apply Nat.zero_le - · apply pow_pos; norm_num - apply add_le_add_left; rw [div_le_div_right (by norm_num)] - rw [le_div_iff₀ (by norm_num), ← mul_pow] - refine le_trans ?_ (le_of_eq (one_pow 3)); apply pow_le_pow_left - · apply le_of_lt; apply mul_pos - · apply div_pos pi_pos; apply pow_pos; norm_num - · apply pow_pos; norm_num - · rw [← le_div_iff₀ (by norm_num)] - refine le_trans ((div_le_div_right ?_).mpr pi_le_four) ?_ - · apply pow_pos; norm_num - · ring_nf; rfl + rw [← div_lt_iff₀ (by norm_num), ← sin_pi_over_two_pow_succ, ← sub_lt_iff_lt_add'] + calc + π / 2 ^ (n + 2) - sin (π / 2 ^ (n + 2)) < (π / 2 ^ (n + 2)) ^ 3 / 4 := + sub_lt_comm.1 <| sin_gt_sub_cube (by positivity) <| div_le_one_of_le₀ ?_ (by positivity) + _ ≤ (4 / 2 ^ (n + 2)) ^ 3 / 4 := by gcongr; exact pi_le_four + _ = 1 / (2 ^ n) ^ 3 / 4 := by simp [add_comm n, pow_add, div_mul_eq_div_div]; norm_num + calc + π ≤ 4 := pi_le_four + _ = 2 ^ (0 + 2) := by norm_num + _ ≤ 2 ^ (n + 2) := by gcongr <;> norm_num refine lt_of_lt_of_le this (le_of_eq ?_); rw [add_mul]; congr 1 · ring simp only [show (4 : ℝ) = 2 ^ 2 by norm_num, ← pow_mul, div_div, ← pow_add] diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean index a4b239c726f35..9c6f753d90a13 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/UniformConvergence.lean @@ -78,9 +78,9 @@ lemma r_lower_bound_on_verticalStrip {A B : ℝ} (h : 0 < B) (hz : z ∈ vertica apply min_le_min hz.2 rw [Real.sqrt_le_sqrt_iff (by apply (r1_pos z).le)] simp only [r1_eq, div_pow, one_div] - rw [inv_le_inv₀ (by positivity) (by positivity), add_le_add_iff_right] - apply div_le_div (sq_nonneg _) _ (by positivity) (pow_le_pow_left h.le hz.2 2) - simpa only [even_two.pow_abs] using pow_le_pow_left (abs_nonneg _) hz.1 2 + rw [inv_le_inv₀ (by positivity) (by positivity), add_le_add_iff_right, ← even_two.pow_abs] + gcongr + exacts [hz.1, hz.2] lemma auxbound1 {c : ℝ} (d : ℝ) (hc : 1 ≤ c ^ 2) : r z ≤ Complex.abs (c * z + d) := by rcases z with ⟨z, hz⟩ diff --git a/Mathlib/NumberTheory/ModularForms/JacobiTheta/Bounds.lean b/Mathlib/NumberTheory/ModularForms/JacobiTheta/Bounds.lean index 5af4baf5ed34a..4e6f7d202b007 100644 --- a/Mathlib/NumberTheory/ModularForms/JacobiTheta/Bounds.lean +++ b/Mathlib/NumberTheory/ModularForms/JacobiTheta/Bounds.lean @@ -99,8 +99,8 @@ lemma summable_f_nat (k : ℕ) (a : ℝ) {t : ℝ} (ht : 0 < t) : Summable (f_na simp_rw [← mul_assoc, f_nat, norm_mul, norm_eq_abs, abs_exp, mul_le_mul_iff_of_pos_right (exp_pos _), ← mul_pow, abs_pow, two_mul] filter_upwards [eventually_ge_atTop (Nat.ceil |a|)] with n hn - apply pow_le_pow_left (abs_nonneg _) ((abs_add_le _ _).trans - (add_le_add (le_of_eq (Nat.abs_cast _)) (Nat.ceil_le.mp hn))) + gcongr + exact (abs_add_le ..).trans (add_le_add (Nat.abs_cast _).le (Nat.ceil_le.mp hn)) section k_eq_zero diff --git a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean index 76ecdd6132820..c735c10548900 100644 --- a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean +++ b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean @@ -223,10 +223,7 @@ theorem seq_next {x : 𝓞 K} (hx : x ≠ 0) : calc _ = ∏ w : InfinitePlace K, w (algebraMap _ K y) ^ mult w := (prod_eq_abs_norm (algebraMap _ K y)).symm - _ ≤ ∏ w : InfinitePlace K, (g w : ℝ) ^ mult w := by - refine prod_le_prod ?_ ?_ - · exact fun _ _ => pow_nonneg (by positivity) _ - · exact fun w _ => pow_le_pow_left (by positivity) (le_of_lt (h_yle w)) (mult w) + _ ≤ ∏ w : InfinitePlace K, (g w : ℝ) ^ mult w := by gcongr with w; exact (h_yle w).le _ ≤ (B : ℝ) := by simp_rw [← NNReal.coe_pow, ← NNReal.coe_prod] exact le_of_eq (congr_arg toReal h_gprod) From d51f67b0be25c12a1451aa94deb43b77bc1aedb2 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 18 Nov 2024 07:12:22 +0000 Subject: [PATCH 049/829] chore: backport some simpler proofs from lean-pr-testing-6053 (#19176) A few things that break under https://github.com/leanprover/lean4/pull/6053 admit nicer proofs anyway, so backport these more robust proofs. --- Mathlib/CategoryTheory/Quotient.lean | 2 +- Mathlib/Data/List/Cycle.lean | 5 +---- Mathlib/Data/TypeVec.lean | 4 +--- 3 files changed, 3 insertions(+), 8 deletions(-) diff --git a/Mathlib/CategoryTheory/Quotient.lean b/Mathlib/CategoryTheory/Quotient.lean index 2379d976800bc..612ea6e8c5677 100644 --- a/Mathlib/CategoryTheory/Quotient.lean +++ b/Mathlib/CategoryTheory/Quotient.lean @@ -174,7 +174,7 @@ theorem lift_unique (Φ : Quotient r ⥤ D) (hΦ : functor r ⋙ Φ = F) : Φ = · rintro _ _ f dsimp [lift, Functor] refine Quot.inductionOn f (fun _ ↦ ?_) -- Porting note: this line was originally an `apply` - simp only [Quot.liftOn_mk, Functor.comp_map] + simp only [heq_eq_eq] congr lemma lift_unique' (F₁ F₂ : Quotient r ⥤ D) (h : functor r ⋙ F₁ = functor r ⋙ F₂) : diff --git a/Mathlib/Data/List/Cycle.lean b/Mathlib/Data/List/Cycle.lean index d0dd29cad6f7f..167f0c1116671 100644 --- a/Mathlib/Data/List/Cycle.lean +++ b/Mathlib/Data/List/Cycle.lean @@ -825,10 +825,7 @@ theorem chain_map {β : Type*} {r : α → α → Prop} (f : β → α) {s : Cyc Quotient.inductionOn s fun l => by cases' l with a l · rfl - dsimp only [Chain, Quotient.liftOn_mk, Cycle.map, Quotient.map', Quot.map, - Quotient.liftOn', Quotient.liftOn, Quot.liftOn_mk, List.map] - rw [← concat_eq_append, ← List.map_concat, List.chain_map f] - simp + · simp [← concat_eq_append, ← List.map_concat, List.chain_map f] nonrec theorem chain_range_succ (r : ℕ → ℕ → Prop) (n : ℕ) : Chain r (List.range n.succ) ↔ r n 0 ∧ ∀ m < n, r m m.succ := by diff --git a/Mathlib/Data/TypeVec.lean b/Mathlib/Data/TypeVec.lean index df15092ea20be..5c337fc5f0cd1 100644 --- a/Mathlib/Data/TypeVec.lean +++ b/Mathlib/Data/TypeVec.lean @@ -604,9 +604,7 @@ theorem dropFun_of_subtype {α} (p : α ⟹ «repeat» (n + 1) Prop) : @[simp] theorem lastFun_of_subtype {α} (p : α ⟹ «repeat» (n + 1) Prop) : - lastFun (ofSubtype p) = _root_.id := by - ext i : 2 - induction i; simp [dropFun, *]; rfl + lastFun (ofSubtype p) = _root_.id := rfl @[simp] theorem dropFun_RelLast' {α : TypeVec n} {β} (R : β → β → Prop) : From 9509b8dc6a58e59be621e34b2936212269d7b3c5 Mon Sep 17 00:00:00 2001 From: Tian Chen <27919714+peakpoint@users.noreply.github.com> Date: Mon, 18 Nov 2024 07:41:43 +0000 Subject: [PATCH 050/829] feat: determinant of reflection (#19016) --- .../InnerProductSpace/Projection.lean | 30 +++++++++++++++++++ 1 file changed, 30 insertions(+) diff --git a/Mathlib/Analysis/InnerProductSpace/Projection.lean b/Mathlib/Analysis/InnerProductSpace/Projection.lean index 953e88b6feb07..c2ea3c5347107 100644 --- a/Mathlib/Analysis/InnerProductSpace/Projection.lean +++ b/Mathlib/Analysis/InnerProductSpace/Projection.lean @@ -978,6 +978,36 @@ theorem reflection_sub {v w : F} (h : ‖v‖ = ‖w‖) : reflection (ℝ ∙ ( variable (K) +section FiniteDimensional + +open Module + +variable [FiniteDimensional 𝕜 K] + +@[simp] +theorem det_reflection : LinearMap.det (reflection K).toLinearMap = (-1) ^ finrank 𝕜 Kᗮ := by + by_cases hK : FiniteDimensional 𝕜 Kᗮ + swap + · rw [finrank_of_infinite_dimensional hK, pow_zero, LinearMap.det_eq_one_of_finrank_eq_zero] + exact finrank_of_infinite_dimensional fun h ↦ hK (h.finiteDimensional_submodule _) + let e := K.prodEquivOfIsCompl _ K.isCompl_orthogonal_of_completeSpace + let b := (finBasis 𝕜 K).prod (finBasis 𝕜 Kᗮ) + have : LinearMap.toMatrix b b (e.symm ∘ₗ (reflection K).toLinearMap ∘ₗ e.symm.symm) = + Matrix.fromBlocks 1 0 0 (-1) := by + ext (_ | _) (_ | _) <;> + simp [LinearMap.toMatrix_apply, b, Matrix.one_apply, Finsupp.single_apply, e, eq_comm, + reflection_mem_subspace_eq_self, reflection_mem_subspace_orthogonalComplement_eq_neg] + rw [← LinearMap.det_conj _ e.symm, ← LinearMap.det_toMatrix b, this, Matrix.det_fromBlocks_zero₂₁, + Matrix.det_one, one_mul, Matrix.det_neg, Fintype.card_fin, Matrix.det_one, mul_one] + +@[simp] +theorem linearEquiv_det_reflection : (reflection K).det = (-1) ^ finrank 𝕜 Kᗮ := by + ext + rw [LinearEquiv.coe_det, Units.val_pow_eq_pow_val] + exact det_reflection K + +end FiniteDimensional + -- Porting note: relax assumptions, swap LHS with RHS /-- If the orthogonal projection to `K` is well-defined, then a vector splits as the sum of its orthogonal projections onto a complete submodule `K` and onto the orthogonal complement of `K`. -/ From 6d297a4172e6c37d3bf82e68924c45d72621ac5d Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Mon, 18 Nov 2024 07:41:44 +0000 Subject: [PATCH 051/829] feat(NumberTheory/LSeries/Nonvanishing): new file (#19043) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This is a substantial step on the way to PNT and Dirichlet's Theorem: it adds the fact that the L-function of a Dirichlet character does not vanish on the closed right half-plane `Re(s) ≥ 1`. From [EulerProducts](https://github.com/MichaelStollBayreuth/EulerProducts). See [here](https://leanprover.zulipchat.com/#narrow/channel/144837-PR-reviews/topic/Prerequisites.20for.20PNT.20and.20Dirichlet's.20Thm/near/482385971) on Zulip. Co-authored-by: Michael Stoll <99838730+MichaelStollBayreuth@users.noreply.github.com> --- Mathlib.lean | 1 + .../NumberTheory/LSeries/Nonvanishing.lean | 238 ++++++++++++++++++ 2 files changed, 239 insertions(+) create mode 100644 Mathlib/NumberTheory/LSeries/Nonvanishing.lean diff --git a/Mathlib.lean b/Mathlib.lean index 3b1c7dbea4b29..9dbb62dd6373b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3752,6 +3752,7 @@ import Mathlib.NumberTheory.LSeries.HurwitzZetaOdd import Mathlib.NumberTheory.LSeries.HurwitzZetaValues import Mathlib.NumberTheory.LSeries.Linearity import Mathlib.NumberTheory.LSeries.MellinEqDirichlet +import Mathlib.NumberTheory.LSeries.Nonvanishing import Mathlib.NumberTheory.LSeries.Positivity import Mathlib.NumberTheory.LSeries.QuadraticNonvanishing import Mathlib.NumberTheory.LSeries.RiemannZeta diff --git a/Mathlib/NumberTheory/LSeries/Nonvanishing.lean b/Mathlib/NumberTheory/LSeries/Nonvanishing.lean new file mode 100644 index 0000000000000..f3b7eb0bc0832 --- /dev/null +++ b/Mathlib/NumberTheory/LSeries/Nonvanishing.lean @@ -0,0 +1,238 @@ +/- +Copyright (c) 2024 Michael Stoll. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Michael Stoll +-/ +import Mathlib.Analysis.SpecialFunctions.Complex.LogBounds +import Mathlib.NumberTheory.Harmonic.ZetaAsymp +import Mathlib.NumberTheory.LSeries.QuadraticNonvanishing + +/-! +# The L-function of a Dirichlet character does not vanish on Re(s) ≥ 1 + +The main result in this file is `DirichletCharacter.Lfunction_ne_zero_of_one_le_re`: +if `χ` is a Dirichlet character and `s.re ≥ 1` and either `χ` is nontrivial or `s ≠ 1`, +then the L-function of `χ` does not vanish at `s`. + +As a consequence, we have the corresponding statement for the Riemann ζ function: +`riemannZeta_ne_zero_of_one_le_re`. + +These results are prerequisites for the **Prime Number Theorem** and +**Dirichlet's Theorem** on primes in arithmetic progressions. +-/ + +section nonvanishing + +open Complex + +-- This is the key positivity lemma that is used to show that the L-function +-- of a Dirichlet character `χ` does not vanish for `s.re ≥ 1` (unless `χ^2 = 1` and `s = 1`). +private lemma re_log_comb_nonneg {a : ℝ} (ha₀ : 0 ≤ a) (ha₁ : a < 1) {z : ℂ} (hz : ‖z‖ = 1) : + 0 ≤ 3 * (-log (1 - a)).re + 4 * (-log (1 - a * z)).re + (-log (1 - a * z ^ 2)).re := by + have hac₀ : ‖(a : ℂ)‖ < 1 := by + simp only [norm_eq_abs, abs_ofReal, _root_.abs_of_nonneg ha₀, ha₁] + have hac₁ : ‖a * z‖ < 1 := by rwa [norm_mul, hz, mul_one] + have hac₂ : ‖a * z ^ 2‖ < 1 := by rwa [norm_mul, norm_pow, hz, one_pow, mul_one] + rw [← ((hasSum_re <| hasSum_taylorSeries_neg_log hac₀).mul_left 3).add + ((hasSum_re <| hasSum_taylorSeries_neg_log hac₁).mul_left 4) |>.add + (hasSum_re <| hasSum_taylorSeries_neg_log hac₂) |>.tsum_eq] + refine tsum_nonneg fun n ↦ ?_ + simp only [← ofReal_pow, div_natCast_re, ofReal_re, mul_pow, mul_re, ofReal_im, zero_mul, + sub_zero] + rcases n.eq_zero_or_pos with rfl | hn + · simp only [pow_zero, Nat.cast_zero, div_zero, mul_zero, one_re, mul_one, add_zero, le_refl] + · simp only [← mul_div_assoc, ← add_div] + refine div_nonneg ?_ n.cast_nonneg + rw [← pow_mul, pow_mul', sq, mul_re, ← sq, ← sq, ← sq_abs_sub_sq_re, ← norm_eq_abs, norm_pow, + hz] + convert (show 0 ≤ 2 * a ^ n * ((z ^ n).re + 1) ^ 2 by positivity) using 1 + ring + +namespace DirichletCharacter + +variable {N : ℕ} (χ : DirichletCharacter ℂ N) + +-- This is the version of the technical positivity lemma for logarithms of Euler factors. +private lemma re_log_comb_nonneg {n : ℕ} (hn : 2 ≤ n) {x : ℝ} (hx : 1 < x) (y : ℝ) : + 0 ≤ 3 * (-log (1 - (1 : DirichletCharacter ℂ N) n * n ^ (-x : ℂ))).re + + 4 * (-log (1 - χ n * n ^ (-(x + I * y)))).re + + (-log (1 - (χ n ^ 2) * n ^ (-(x + 2 * I * y)))).re := by + by_cases hn' : IsUnit (n : ZMod N) + · have ha₀ : 0 ≤ (n : ℝ) ^ (-x) := Real.rpow_nonneg n.cast_nonneg _ + have ha₁ : (n : ℝ) ^ (-x) < 1 := by + rw [Real.rpow_neg (Nat.cast_nonneg n), inv_lt_one_iff₀] + exact .inr <| Real.one_lt_rpow (mod_cast one_lt_two.trans_le hn) <| zero_lt_one.trans hx + have hz : ‖χ n * (n : ℂ) ^ (-(I * y))‖ = 1 := by + rw [norm_mul, ← hn'.unit_spec, DirichletCharacter.unit_norm_eq_one χ hn'.unit, + norm_eq_abs, ← ofReal_natCast, abs_cpow_eq_rpow_re_of_pos (mod_cast by omega)] + simp only [neg_re, mul_re, I_re, ofReal_re, zero_mul, I_im, ofReal_im, mul_zero, sub_self, + neg_zero, Real.rpow_zero, one_mul] + rw [MulChar.one_apply hn', one_mul] + convert _root_.re_log_comb_nonneg ha₀ ha₁ hz using 6 + · simp only [ofReal_cpow n.cast_nonneg (-x), ofReal_natCast, ofReal_neg] + · congr 2 + rw [neg_add, cpow_add _ _ <| mod_cast by omega, ← ofReal_neg, ofReal_cpow n.cast_nonneg (-x), + ofReal_natCast, mul_left_comm] + · rw [neg_add, cpow_add _ _ <| mod_cast by omega, ← ofReal_neg, ofReal_cpow n.cast_nonneg (-x), + ofReal_natCast, show -(2 * I * y) = (2 : ℕ) * -(I * y) by ring, cpow_nat_mul, mul_pow, + mul_left_comm] + · simp only [MulChar.map_nonunit _ hn', zero_mul, sub_zero, log_one, neg_zero, zero_re, mul_zero, + neg_add_rev, add_zero, pow_two, le_refl] + +/-- The logarithms of the Euler factors of a Dirichlet L-series form a summable sequence. -/ +lemma summable_neg_log_one_sub_mul_prime_cpow {s : ℂ} (hs : 1 < s.re) : + Summable fun p : Nat.Primes ↦ -log (1 - χ p * (p : ℂ) ^ (-s)) := by + have (p : Nat.Primes) : ‖χ p * (p : ℂ) ^ (-s)‖ ≤ (p : ℝ) ^ (-s).re := by + simpa only [norm_mul, norm_natCast_cpow_of_re_ne_zero _ <| re_neg_ne_zero_of_one_lt_re hs] + using mul_le_of_le_one_left (by positivity) (χ.norm_le_one _) + refine (Nat.Primes.summable_rpow.mpr ?_).of_nonneg_of_le (fun _ ↦ norm_nonneg _) this + |>.of_norm.clog_one_sub.neg + simp only [neg_re, neg_lt_neg_iff, hs] + +private lemma one_lt_re_one_add {x : ℝ} (hx : 0 < x) (y : ℝ) : + 1 < (1 + x : ℂ).re ∧ 1 < (1 + x + I * y).re ∧ 1 < (1 + x + 2 * I * y).re := by + simp only [add_re, one_re, ofReal_re, lt_add_iff_pos_right, hx, mul_re, I_re, zero_mul, I_im, + ofReal_im, mul_zero, sub_self, add_zero, re_ofNat, im_ofNat, mul_one, mul_im, and_self] + +open scoped LSeries.notation in +/-- For positive `x` and nonzero `y` and a Dirichlet character `χ` we have that +`|L(χ^0, x)^3 L(χ, x+iy)^4 L(χ^2, x+2iy)| ≥ 1. -/ +lemma norm_LSeries_product_ge_one {x : ℝ} (hx : 0 < x) (y : ℝ) : + ‖L ↗(1 : DirichletCharacter ℂ N) (1 + x) ^ 3 * L ↗χ (1 + x + I * y) ^ 4 * + L ↗(χ ^ 2 :) (1 + x + 2 * I * y)‖ ≥ 1 := by + have ⟨h₀, h₁, h₂⟩ := one_lt_re_one_add hx y + have H₀ := summable_neg_log_one_sub_mul_prime_cpow (N := N) 1 h₀ + have H₁ := summable_neg_log_one_sub_mul_prime_cpow χ h₁ + have H₂ := summable_neg_log_one_sub_mul_prime_cpow (χ ^ 2) h₂ + have hsum₀ := (hasSum_re H₀.hasSum).summable.mul_left 3 + have hsum₁ := (hasSum_re H₁.hasSum).summable.mul_left 4 + have hsum₂ := (hasSum_re H₂.hasSum).summable + rw [← LSeries_eulerProduct_exp_log _ h₀, ← LSeries_eulerProduct_exp_log χ h₁, + ← LSeries_eulerProduct_exp_log _ h₂] + simp only [← exp_nat_mul, Nat.cast_ofNat, ← exp_add, norm_eq_abs, abs_exp, add_re, mul_re, + re_ofNat, im_ofNat, zero_mul, sub_zero, Real.one_le_exp_iff] + rw [re_tsum H₀, re_tsum H₁, re_tsum H₂, ← tsum_mul_left, ← tsum_mul_left, + ← tsum_add hsum₀ hsum₁, ← tsum_add (hsum₀.add hsum₁) hsum₂] + simpa only [neg_add_rev, neg_re, mul_neg, χ.pow_apply' two_ne_zero, ge_iff_le, add_re, one_re, + ofReal_re, ofReal_add, ofReal_one] using + tsum_nonneg fun (p : Nat.Primes) ↦ χ.re_log_comb_nonneg p.prop.two_le h₀ y + +variable [NeZero N] + +/-- A variant of `DirichletCharacter.norm_LSeries_product_ge_one` in terms of the L-functions. -/ +lemma norm_LFunction_product_ge_one {x : ℝ} (hx : 0 < x) (y : ℝ) : + ‖LFunctionTrivChar N (1 + x) ^ 3 * LFunction χ (1 + x + I * y) ^ 4 * + LFunction (χ ^ 2) (1 + x + 2 * I * y)‖ ≥ 1 := by + have ⟨h₀, h₁, h₂⟩ := one_lt_re_one_add hx y + rw [LFunctionTrivChar, DirichletCharacter.LFunction_eq_LSeries 1 h₀, + χ.LFunction_eq_LSeries h₁, (χ ^ 2).LFunction_eq_LSeries h₂] + exact norm_LSeries_product_ge_one χ hx y + +open Asymptotics Topology Filter + +open Homeomorph in +lemma LFunctionTrivChar_isBigO_near_one_horizontal : + (fun x : ℝ ↦ LFunctionTrivChar N (1 + x)) =O[𝓝[>] 0] fun x ↦ (1 : ℂ) / x := by + have : (fun w : ℂ ↦ LFunctionTrivChar N (1 + w)) =O[𝓝[≠] 0] (1 / ·) := by + have H : Tendsto (fun w ↦ w * LFunctionTrivChar N (1 + w)) (𝓝[≠] 0) + (𝓝 <| ∏ p ∈ N.primeFactors, (1 - (p : ℂ)⁻¹)) := by + convert (LFunctionTrivChar_residue_one (N := N)).comp (f := fun w ↦ 1 + w) ?_ using 1 + · simp only [Function.comp_def, add_sub_cancel_left] + · simpa only [tendsto_iff_comap, Homeomorph.coe_addLeft, add_zero, map_le_iff_le_comap] using + ((Homeomorph.addLeft (1 : ℂ)).map_punctured_nhds_eq 0).le + exact (isBigO_mul_iff_isBigO_div eventually_mem_nhdsWithin).mp <| H.isBigO_one ℂ + exact (isBigO_comp_ofReal_nhds_ne this).mono <| nhds_right'_le_nhds_ne 0 + +omit [NeZero N] in +private lemma one_add_I_mul_ne_one_or {y : ℝ} (hy : y ≠ 0 ∨ χ ≠ 1) : + 1 + I * y ≠ 1 ∨ χ ≠ 1:= by + simpa only [ne_eq, add_right_eq_self, _root_.mul_eq_zero, I_ne_zero, ofReal_eq_zero, false_or] + using hy + +lemma LFunction_isBigO_horizontal {y : ℝ} (hy : y ≠ 0 ∨ χ ≠ 1) : + (fun x : ℝ ↦ LFunction χ (1 + x + I * y)) =O[𝓝[>] 0] fun _ ↦ (1 : ℂ) := by + refine IsBigO.mono ?_ nhdsWithin_le_nhds + simp_rw [add_comm (1 : ℂ), add_assoc] + have := (χ.differentiableAt_LFunction _ <| one_add_I_mul_ne_one_or χ hy).continuousAt + rw [← zero_add (1 + _)] at this + exact this.comp (f := fun x : ℝ ↦ x + (1 + I * y)) (x := 0) (by fun_prop) |>.tendsto.isBigO_one ℂ + +private lemma LFunction_isBigO_horizontal_of_eq_zero {y : ℝ} (hy : y ≠ 0 ∨ χ ≠ 1) + (h : LFunction χ (1 + I * y) = 0) : + (fun x : ℝ ↦ LFunction χ (1 + x + I * y)) =O[𝓝[>] 0] fun x : ℝ ↦ (x : ℂ) := by + simp_rw [add_comm (1 : ℂ), add_assoc] + have := (χ.differentiableAt_LFunction _ <| one_add_I_mul_ne_one_or χ hy).hasDerivAt + rw [← zero_add (1 + _)] at this + simpa only [zero_add, h, sub_zero] + using (Complex.isBigO_comp_ofReal_nhds + (this.comp_add_const 0 _).differentiableAt.isBigO_sub) |>.mono nhdsWithin_le_nhds + +-- intermediate statement, special case of the next theorem +private lemma LFunction_ne_zero_of_not_quadratic_or_ne_one {t : ℝ} (h : χ ^ 2 ≠ 1 ∨ t ≠ 0) : + LFunction χ (1 + I * t) ≠ 0 := by + intro Hz + have hz₁ : t ≠ 0 ∨ χ ≠ 1 := by + refine h.symm.imp_right (fun h H ↦ ?_) + simp only [H, one_pow, ne_eq, not_true_eq_false] at h + have hz₂ : 2 * t ≠ 0 ∨ χ ^ 2 ≠ 1 := + h.symm.imp_left <| mul_ne_zero two_ne_zero + have help (x : ℝ) : ((1 / x) ^ 3 * x ^ 4 * 1 : ℂ) = x := by + rcases eq_or_ne x 0 with rfl | h + · rw [ofReal_zero, zero_pow (by omega), mul_zero, mul_one] + · rw [one_div, inv_pow, pow_succ _ 3, ← mul_assoc, + inv_mul_cancel₀ <| pow_ne_zero 3 (ofReal_ne_zero.mpr h), one_mul, mul_one] + -- put together the various `IsBigO` statements and `norm_LFunction_product_ge_one` + -- to derive a contradiction + have H₀ : (fun _ : ℝ ↦ (1 : ℝ)) =O[𝓝[>] 0] + fun x ↦ LFunctionTrivChar N (1 + x) ^ 3 * LFunction χ (1 + x + I * t) ^ 4 * + LFunction (χ ^ 2) (1 + x + 2 * I * t) := + IsBigO.of_bound' <| eventually_nhdsWithin_of_forall + fun _ hx ↦ (norm_one (α := ℝ)).symm ▸ (χ.norm_LFunction_product_ge_one hx t).le + have H := (LFunctionTrivChar_isBigO_near_one_horizontal (N := N)).pow 3 |>.mul <| + (χ.LFunction_isBigO_horizontal_of_eq_zero hz₁ Hz).pow 4 |>.mul <| + LFunction_isBigO_horizontal _ hz₂ + simp only [ofReal_mul, ofReal_ofNat, mul_left_comm I, ← mul_assoc, help] at H + -- go via absolute value to translate into a statement over `ℝ` + replace H := (H₀.trans H).norm_right + simp only [norm_eq_abs, abs_ofReal] at H + exact isLittleO_irrefl (.of_forall (fun _ ↦ one_ne_zero)) <| H.of_norm_right.trans_isLittleO + <| isLittleO_id_one.mono nhdsWithin_le_nhds + +/-- If `χ` is a Dirichlet character, then `L(χ, s)` does not vanish when `s.re = 1` +except when `χ` is trivial and `s = 1` (then `L(χ, s)` has a simple pole at `s = 1`). -/ +theorem LFunction_ne_zero_of_re_eq_one {s : ℂ} (hs : s.re = 1) (hχs : χ ≠ 1 ∨ s ≠ 1) : + LFunction χ s ≠ 0 := by + by_cases h : χ ^ 2 = 1 ∧ s = 1 + · exact h.2 ▸ LFunction_at_one_ne_zero_of_quadratic h.1 <| hχs.neg_resolve_right h.2 + · have hs' : s = 1 + I * s.im := by + conv_lhs => rw [← re_add_im s, hs, ofReal_one, mul_comm] + rw [not_and_or, ← ne_eq, ← ne_eq, hs', add_right_ne_self] at h + replace h : χ ^ 2 ≠ 1 ∨ s.im ≠ 0 := + h.imp_right (fun H ↦ by exact_mod_cast right_ne_zero_of_mul H) + exact hs'.symm ▸ χ.LFunction_ne_zero_of_not_quadratic_or_ne_one h + +/-- If `χ` is a Dirichlet character, then `L(χ, s)` does not vanish for `s.re ≥ 1` +except when `χ` is trivial and `s = 1` (then `L(χ, s)` has a simple pole at `s = 1`). -/ +theorem LFunction_ne_zero_of_one_le_re ⦃s : ℂ⦄ (hχs : χ ≠ 1 ∨ s ≠ 1) (hs : 1 ≤ s.re) : + LFunction χ s ≠ 0 := + hs.eq_or_lt.casesOn (fun hs ↦ LFunction_ne_zero_of_re_eq_one χ hs.symm hχs) + fun hs ↦ LFunction_eq_LSeries χ hs ▸ LSeries_ne_zero_of_one_lt_re χ hs + +-- Interesting special case: +variable {χ} in +/-- The L-function of a nontrivial Dirichlet character does not vanish at `s = 1`. -/ +theorem LFunction_apply_one_ne_zero (hχ : χ ≠ 1) : LFunction χ 1 ≠ 0 := + LFunction_ne_zero_of_one_le_re χ (.inl hχ) <| one_re ▸ le_rfl + +end DirichletCharacter + +open DirichletCharacter in +/-- The Riemann Zeta Function does not vanish on the closed half-plane `re s ≥ 1`. +(Note that the value at `s = 1` is a junk value, which happens to be nonzero.) -/ +lemma riemannZeta_ne_zero_of_one_le_re ⦃s : ℂ⦄ (hs : 1 ≤ s.re) : + riemannZeta s ≠ 0 := by + rcases eq_or_ne s 1 with rfl | hs₀ + · exact riemannZeta_one_ne_zero + · exact LFunction_modOne_eq (χ := 1) ▸ LFunction_ne_zero_of_one_le_re _ (.inr hs₀) hs + +end nonvanishing From 66df24a8e2b60b44dbe1def58ad03a083b80e865 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Mon, 18 Nov 2024 09:35:44 +0000 Subject: [PATCH 052/829] refactor(GroupTheory/SpecificGroups/Cyclic): Switch from `Fintype` to `Finite` (#19109) This PR switches most of `GroupTheory/SpecificGroups/Cyclic.lean` from `Fintype` to `Finite`. --- Mathlib/FieldTheory/Finite/Basic.lean | 10 +- Mathlib/GroupTheory/Exponent.lean | 2 + .../GroupTheory/SpecificGroups/Cyclic.lean | 166 ++++++++---------- .../SpecificGroups/Quaternion.lean | 2 +- .../Cyclotomic/CyclotomicCharacter.lean | 2 +- Mathlib/NumberTheory/LucasPrimality.lean | 3 +- Mathlib/NumberTheory/MulChar/Lemmas.lean | 2 +- Mathlib/RingTheory/RootsOfUnity/Basic.lean | 11 +- .../RootsOfUnity/EnoughRootsOfUnity.lean | 3 +- 9 files changed, 92 insertions(+), 109 deletions(-) diff --git a/Mathlib/FieldTheory/Finite/Basic.lean b/Mathlib/FieldTheory/Finite/Basic.lean index b7a753f2dc315..59645c6b82ad7 100644 --- a/Mathlib/FieldTheory/Finite/Basic.lean +++ b/Mathlib/FieldTheory/Finite/Basic.lean @@ -172,6 +172,7 @@ theorem sum_subgroup_units [Ring K] [NoZeroDivisors K] theorem sum_subgroup_pow_eq_zero [CommRing K] [NoZeroDivisors K] {G : Subgroup Kˣ} [Fintype G] {k : ℕ} (k_pos : k ≠ 0) (k_lt_card_G : k < Fintype.card G) : ∑ x : G, ((x : Kˣ) : K) ^ k = 0 := by + rw [← Nat.card_eq_fintype_card] at k_lt_card_G nontriviality K have := NoZeroDivisors.to_isDomain K rcases (exists_pow_ne_one_of_isCyclic k_pos k_lt_card_G) with ⟨a, ha⟩ @@ -255,7 +256,7 @@ theorem cast_card_eq_zero : (q : K) = 0 := by theorem forall_pow_eq_one_iff (i : ℕ) : (∀ x : Kˣ, x ^ i = 1) ↔ q - 1 ∣ i := by classical obtain ⟨x, hx⟩ := IsCyclic.exists_generator (α := Kˣ) - rw [← Fintype.card_units, ← orderOf_eq_card_of_forall_mem_zpowers hx, + rw [← Nat.card_eq_fintype_card, ← Nat.card_units, ← orderOf_eq_card_of_forall_mem_zpowers hx, orderOf_dvd_iff_pow_eq_one] constructor · intro h; apply h @@ -586,11 +587,12 @@ theorem unit_isSquare_iff (hF : ringChar F ≠ 2) (a : Fˣ) : push_cast exact FiniteField.pow_card_sub_one_eq_one (y : F) (Units.ne_zero y) · subst a; intro h - have key : 2 * (Fintype.card F / 2) ∣ n * (Fintype.card F / 2) := by + rw [← Nat.card_eq_fintype_card] at hodd h + have key : 2 * (Nat.card F / 2) ∣ n * (Nat.card F / 2) := by rw [← pow_mul] at h - rw [hodd, ← Fintype.card_units, ← orderOf_eq_card_of_forall_mem_zpowers hg] + rw [hodd, ← Nat.card_units, ← orderOf_eq_card_of_forall_mem_zpowers hg] apply orderOf_dvd_of_pow_eq_one h - have : 0 < Fintype.card F / 2 := Nat.div_pos Fintype.one_lt_card (by norm_num) + have : 0 < Nat.card F / 2 := Nat.div_pos Finite.one_lt_card (by norm_num) obtain ⟨m, rfl⟩ := Nat.dvd_of_mul_dvd_mul_right this key refine ⟨g ^ m, ?_⟩ dsimp diff --git a/Mathlib/GroupTheory/Exponent.lean b/Mathlib/GroupTheory/Exponent.lean index 078d6182ec058..de0f7edebeedc 100644 --- a/Mathlib/GroupTheory/Exponent.lean +++ b/Mathlib/GroupTheory/Exponent.lean @@ -501,9 +501,11 @@ variable [Group G] lemma Group.one_lt_exponent [Finite G] [Nontrivial G] : 1 < Monoid.exponent G := Monoid.one_lt_exponent +@[to_additive] theorem Group.exponent_dvd_card [Fintype G] : Monoid.exponent G ∣ Fintype.card G := Monoid.exponent_dvd.mpr <| fun _ => orderOf_dvd_card +@[to_additive] theorem Group.exponent_dvd_nat_card : Monoid.exponent G ∣ Nat.card G := Monoid.exponent_dvd.mpr orderOf_dvd_natCard diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index cdf9bb10221cb..ca2a56cc761c7 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -100,65 +100,65 @@ theorem MonoidHom.map_cyclic [h : IsCyclic G] (σ : G →* G) : MonoidAddHom.map_add_cyclic := AddMonoidHom.map_addCyclic @[to_additive] -theorem isCyclic_of_orderOf_eq_card [Fintype α] (x : α) (hx : orderOf x = Fintype.card α) : +theorem isCyclic_of_orderOf_eq_card [Finite α] (x : α) (hx : orderOf x = Nat.card α) : IsCyclic α := by + cases nonempty_fintype α use x rw [← Set.range_eq_univ, ← coe_zpowers] - rw [← Fintype.card_congr (Equiv.Set.univ α), ← Fintype.card_zpowers] at hx + rw [← Nat.card_congr (Equiv.Set.univ α), Nat.card_eq_fintype_card, ← Fintype.card_zpowers] at hx convert Set.eq_of_subset_of_card_le (Set.subset_univ _) (ge_of_eq hx) @[deprecated (since := "2024-02-21")] alias isAddCyclic_of_orderOf_eq_card := isAddCyclic_of_addOrderOf_eq_card @[to_additive] -theorem Subgroup.eq_bot_or_eq_top_of_prime_card {_ : Fintype G} - (H : Subgroup G) [hp : Fact (Fintype.card G).Prime] : H = ⊥ ∨ H = ⊤ := by - classical +theorem Subgroup.eq_bot_or_eq_top_of_prime_card + (H : Subgroup G) [hp : Fact (Nat.card G).Prime] : H = ⊥ ∨ H = ⊤ := by + have : Finite G := Nat.finite_of_card_ne_zero hp.1.ne_zero have := card_subgroup_dvd_card H - rwa [Nat.card_eq_fintype_card (α := G), Nat.dvd_prime hp.1, ← Nat.card_eq_fintype_card, - ← eq_bot_iff_card, card_eq_iff_eq_top] at this + rwa [Nat.dvd_prime hp.1, ← eq_bot_iff_card, card_eq_iff_eq_top] at this /-- Any non-identity element of a finite group of prime order generates the group. -/ @[to_additive "Any non-identity element of a finite group of prime order generates the group."] -theorem zpowers_eq_top_of_prime_card {_ : Fintype G} {p : ℕ} - [hp : Fact p.Prime] (h : Fintype.card G = p) {g : G} (hg : g ≠ 1) : zpowers g = ⊤ := by +theorem zpowers_eq_top_of_prime_card {p : ℕ} + [hp : Fact p.Prime] (h : Nat.card G = p) {g : G} (hg : g ≠ 1) : zpowers g = ⊤ := by subst h have := (zpowers g).eq_bot_or_eq_top_of_prime_card rwa [zpowers_eq_bot, or_iff_right hg] at this @[to_additive] -theorem mem_zpowers_of_prime_card {_ : Fintype G} {p : ℕ} [hp : Fact p.Prime] - (h : Fintype.card G = p) {g g' : G} (hg : g ≠ 1) : g' ∈ zpowers g := by +theorem mem_zpowers_of_prime_card {p : ℕ} [hp : Fact p.Prime] + (h : Nat.card G = p) {g g' : G} (hg : g ≠ 1) : g' ∈ zpowers g := by simp_rw [zpowers_eq_top_of_prime_card h hg, Subgroup.mem_top] @[to_additive] -theorem mem_powers_of_prime_card {_ : Fintype G} {p : ℕ} [hp : Fact p.Prime] - (h : Fintype.card G = p) {g g' : G} (hg : g ≠ 1) : g' ∈ Submonoid.powers g := by +theorem mem_powers_of_prime_card {p : ℕ} [hp : Fact p.Prime] + (h : Nat.card G = p) {g g' : G} (hg : g ≠ 1) : g' ∈ Submonoid.powers g := by + have : Finite G := Nat.finite_of_card_ne_zero (h ▸ hp.1.ne_zero) rw [mem_powers_iff_mem_zpowers] exact mem_zpowers_of_prime_card h hg @[to_additive] -theorem powers_eq_top_of_prime_card {_ : Fintype G} {p : ℕ} - [hp : Fact p.Prime] (h : Fintype.card G = p) {g : G} (hg : g ≠ 1) : Submonoid.powers g = ⊤ := by +theorem powers_eq_top_of_prime_card {p : ℕ} + [hp : Fact p.Prime] (h : Nat.card G = p) {g : G} (hg : g ≠ 1) : Submonoid.powers g = ⊤ := by ext x simp [mem_powers_of_prime_card h hg] /-- A finite group of prime order is cyclic. -/ @[to_additive "A finite group of prime order is cyclic."] -theorem isCyclic_of_prime_card [Fintype α] {p : ℕ} [hp : Fact p.Prime] - (h : Fintype.card α = p) : IsCyclic α := by - obtain ⟨g, hg⟩ : ∃ g, g ≠ 1 := Fintype.exists_ne_of_one_lt_card (h.symm ▸ hp.1.one_lt) 1 +theorem isCyclic_of_prime_card {p : ℕ} [hp : Fact p.Prime] + (h : Nat.card α = p) : IsCyclic α := by + have : Finite α := Nat.finite_of_card_ne_zero (h ▸ hp.1.ne_zero) + have : Nontrivial α := Finite.one_lt_card_iff_nontrivial.mp (h ▸ hp.1.one_lt) + obtain ⟨g, hg⟩ : ∃ g : α, g ≠ 1 := exists_ne 1 exact ⟨g, fun g' ↦ mem_zpowers_of_prime_card h hg⟩ /-- A finite group of order dividing a prime is cyclic. -/ @[to_additive "A finite group of order dividing a prime is cyclic."] theorem isCyclic_of_card_dvd_prime {p : ℕ} [hp : Fact p.Prime] (h : Nat.card α ∣ p) : IsCyclic α := by - have : Finite α := Nat.finite_of_card_ne_zero (ne_zero_of_dvd_ne_zero hp.1.ne_zero h) rcases (Nat.dvd_prime hp.out).mp h with h | h · exact @isCyclic_of_subsingleton α _ (Nat.card_eq_one_iff_unique.mp h).1 - · have : Fintype α := Fintype.ofFinite α - rw [Nat.card_eq_fintype_card] at h - exact isCyclic_of_prime_card h + · exact isCyclic_of_prime_card h @[to_additive] theorem isCyclic_of_surjective {F : Type*} [hH : IsCyclic G'] @@ -171,47 +171,31 @@ theorem isCyclic_of_surjective {F : Type*} [hH : IsCyclic G'] exact ⟨n, (map_zpow _ _ _).symm⟩ @[to_additive] -theorem orderOf_eq_card_of_forall_mem_zpowers [Fintype α] {g : α} (hx : ∀ x, x ∈ zpowers g) : - orderOf g = Fintype.card α := by - classical - rw [← Fintype.card_zpowers] - apply Fintype.card_of_finset' - simpa using hx +theorem orderOf_eq_card_of_forall_mem_zpowers {g : α} (hx : ∀ x, x ∈ zpowers g) : + orderOf g = Nat.card α := by + rw [← Nat.card_zpowers, (zpowers g).eq_top_iff'.mpr hx, card_top] -@[to_additive] -lemma orderOf_generator_eq_natCard (h : ∀ x, x ∈ Subgroup.zpowers a) : orderOf a = Nat.card α := - Nat.card_zpowers a ▸ (Nat.card_congr <| Equiv.subtypeUnivEquiv h) +@[deprecated (since := "2024-11-15")] +alias orderOf_generator_eq_natCard := orderOf_eq_card_of_forall_mem_zpowers + +@[deprecated (since := "2024-11-15")] +alias addOrderOf_generator_eq_natCard := addOrderOf_eq_card_of_forall_mem_zmultiples @[to_additive] -theorem exists_pow_ne_one_of_isCyclic [Fintype G] [G_cyclic : IsCyclic G] - {k : ℕ} (k_pos : k ≠ 0) (k_lt_card_G : k < Fintype.card G) : ∃ a : G, a ^ k ≠ 1 := by +theorem exists_pow_ne_one_of_isCyclic [G_cyclic : IsCyclic G] + {k : ℕ} (k_pos : k ≠ 0) (k_lt_card_G : k < Nat.card G) : ∃ a : G, a ^ k ≠ 1 := by + have : Finite G := Nat.finite_of_card_ne_zero (Nat.not_eq_zero_of_lt k_lt_card_G) rcases G_cyclic with ⟨a, ha⟩ use a contrapose! k_lt_card_G convert orderOf_le_of_pow_eq_one k_pos.bot_lt k_lt_card_G - rw [← Nat.card_eq_fintype_card, ← Nat.card_zpowers, eq_comm, card_eq_iff_eq_top, eq_top_iff] + rw [← Nat.card_zpowers, eq_comm, card_eq_iff_eq_top, eq_top_iff] exact fun x _ ↦ ha x @[to_additive] theorem Infinite.orderOf_eq_zero_of_forall_mem_zpowers [Infinite α] {g : α} (h : ∀ x, x ∈ zpowers g) : orderOf g = 0 := by - classical - rw [orderOf_eq_zero_iff'] - refine fun n hn hgn => ?_ - have ho := isOfFinOrder_iff_pow_eq_one.mpr ⟨n, hn, hgn⟩ - obtain ⟨x, hx⟩ := - Infinite.exists_not_mem_finset - (Finset.image (fun x => g ^ x) <| Finset.range <| orderOf g) - apply hx - rw [← ho.mem_powers_iff_mem_range_orderOf, Submonoid.mem_powers_iff] - obtain ⟨k, hk⟩ := h x - dsimp at hk - obtain ⟨k, rfl | rfl⟩ := k.eq_nat_or_neg - · exact ⟨k, mod_cast hk⟩ - rw [← zpow_mod_orderOf] at hk - have : 0 ≤ (-k % orderOf g : ℤ) := Int.emod_nonneg (-k) (mod_cast ho.orderOf_pos.ne') - refine ⟨(-k % orderOf g : ℤ).toNat, ?_⟩ - rwa [← zpow_natCast, Int.toNat_of_nonneg this] + rw [orderOf_eq_card_of_forall_mem_zpowers h, Nat.card_eq_zero_of_infinite] @[to_additive] instance Bot.isCyclic : IsCyclic (⊥ : Subgroup α) := @@ -300,8 +284,8 @@ theorem IsCyclic.card_pow_eq_one_le [DecidableEq α] [Fintype α] [IsCyclic α] rw [zpow_natCast, ← pow_mul, Nat.mul_div_cancel_left', hm] refine Nat.dvd_of_mul_dvd_mul_right (gcd_pos_of_pos_left (Fintype.card α) hn0) ?_ conv_lhs => - rw [Nat.div_mul_cancel (Nat.gcd_dvd_right _ _), ← - orderOf_eq_card_of_forall_mem_zpowers hg] + rw [Nat.div_mul_cancel (Nat.gcd_dvd_right _ _), ← Nat.card_eq_fintype_card, + ← orderOf_eq_card_of_forall_mem_zpowers hg] exact orderOf_dvd_of_pow_eq_one hgmn⟩ _ ≤ n := by let ⟨m, hm⟩ := Nat.gcd_dvd_right n (Fintype.card α) @@ -310,7 +294,8 @@ theorem IsCyclic.card_pow_eq_one_le [DecidableEq α] [Fintype α] [IsCyclic α] rw [hm0, mul_zero, Fintype.card_eq_zero_iff] at hm exact hm.elim' 1 simp only [Set.toFinset_card, SetLike.coe_sort_coe] - rw [Fintype.card_zpowers, orderOf_pow g, orderOf_eq_card_of_forall_mem_zpowers hg] + rw [Fintype.card_zpowers, orderOf_pow g, orderOf_eq_card_of_forall_mem_zpowers hg, + Nat.card_eq_fintype_card] nth_rw 2 [hm]; nth_rw 3 [hm] rw [Nat.mul_div_cancel_left _ (gcd_pos_of_pos_left _ hn0), gcd_mul_left_left, hm, Nat.mul_div_cancel _ hm0] @@ -356,9 +341,9 @@ theorem IsCyclic.unique_zpow_zmod (ha : ∀ x : α, x ∈ zpowers a) (x : α) : obtain ⟨n, rfl⟩ := ha x refine ⟨n, (?_ : a ^ n = _), fun y (hy : a ^ n = _) ↦ ?_⟩ · rw [← zpow_natCast, zpow_eq_zpow_iff_modEq, orderOf_eq_card_of_forall_mem_zpowers ha, - Int.modEq_comm, Int.modEq_iff_add_fac, ← ZMod.intCast_eq_iff] + Int.modEq_comm, Int.modEq_iff_add_fac, Nat.card_eq_fintype_card, ← ZMod.intCast_eq_iff] · rw [← zpow_natCast, zpow_eq_zpow_iff_modEq, orderOf_eq_card_of_forall_mem_zpowers ha, - ← ZMod.intCast_eq_intCast_iff] at hy + Nat.card_eq_fintype_card, ← ZMod.intCast_eq_intCast_iff] at hy simp [hy] variable [DecidableEq α] @@ -371,12 +356,13 @@ theorem IsCyclic.image_range_orderOf (ha : ∀ x : α, x ∈ zpowers a) : @[to_additive] theorem IsCyclic.image_range_card (ha : ∀ x : α, x ∈ zpowers a) : - Finset.image (fun i => a ^ i) (range (Fintype.card α)) = univ := by + Finset.image (fun i => a ^ i) (range (Nat.card α)) = univ := by rw [← orderOf_eq_card_of_forall_mem_zpowers ha, IsCyclic.image_range_orderOf ha] @[to_additive] -lemma IsCyclic.ext [Fintype G] [IsCyclic G] {d : ℕ} {a b : ZMod d} - (hGcard : Fintype.card G = d) (h : ∀ t : G, t ^ a.val = t ^ b.val) : a = b := by +lemma IsCyclic.ext [Finite G] [IsCyclic G] {d : ℕ} {a b : ZMod d} + (hGcard : Nat.card G = d) (h : ∀ t : G, t ^ a.val = t ^ b.val) : a = b := by + have : NeZero (Nat.card G) := ⟨Nat.card_pos.ne'⟩ obtain ⟨g, hg⟩ := IsCyclic.exists_generator (α := G) specialize h g subst hGcard @@ -474,9 +460,9 @@ theorem card_orderOf_eq_totient_aux₂ {d : ℕ} (hd : d ∣ Fintype.card α) : @[to_additive isAddCyclic_of_card_nsmul_eq_zero_le, stacks 09HX "This theorem is stronger than \ 09HX. It removes the abelian condition, and requires only `≤` instead of `=`."] theorem isCyclic_of_card_pow_eq_one_le : IsCyclic α := - have : Finset.Nonempty {a : α | orderOf a = Fintype.card α} := + have : Finset.Nonempty {a : α | orderOf a = Nat.card α} := card_pos.1 <| by - rw [card_orderOf_eq_totient_aux₂ hn dvd_rfl, totient_pos] + rw [Nat.card_eq_fintype_card, card_orderOf_eq_totient_aux₂ hn dvd_rfl, totient_pos] apply Fintype.card_pos let ⟨x, hx⟩ := this isCyclic_of_orderOf_eq_card x (Finset.mem_filter.1 hx).2 @@ -496,12 +482,11 @@ alias IsAddCyclic.card_orderOf_eq_totient := IsAddCyclic.card_addOrderOf_eq_toti /-- A finite group of prime order is simple. -/ @[to_additive "A finite group of prime order is simple."] -theorem isSimpleGroup_of_prime_card [Fintype α] {p : ℕ} [hp : Fact p.Prime] - (h : Fintype.card α = p) : IsSimpleGroup α := by +theorem isSimpleGroup_of_prime_card {p : ℕ} [hp : Fact p.Prime] + (h : Nat.card α = p) : IsSimpleGroup α := by subst h - have : Nontrivial α := by - have h' := Nat.Prime.one_lt hp.out - exact Fintype.one_lt_card_iff_nontrivial.1 h' + have : Finite α := Nat.finite_of_card_ne_zero hp.1.ne_zero + have : Nontrivial α := Finite.one_lt_card_iff_nontrivial.mp hp.1.one_lt exact ⟨fun H _ => H.eq_bot_or_eq_top_of_prime_card⟩ end Cyclic @@ -563,18 +548,18 @@ instance (priority := 100) isCyclic : IsCyclic α := by exact ⟨⟨g, (Subgroup.eq_top_iff' _).1 this⟩⟩ @[to_additive] -theorem prime_card [Fintype α] : (Fintype.card α).Prime := by - have h0 : 0 < Fintype.card α := Fintype.card_pos_iff.2 (by infer_instance) +theorem prime_card [Finite α] : (Nat.card α).Prime := by + have h0 : 0 < Nat.card α := Nat.card_pos obtain ⟨g, hg⟩ := IsCyclic.exists_generator (α := α) rw [Nat.prime_def_lt''] - refine ⟨Fintype.one_lt_card_iff_nontrivial.2 inferInstance, fun n hn => ?_⟩ + refine ⟨Finite.one_lt_card_iff_nontrivial.2 inferInstance, fun n hn => ?_⟩ refine (IsSimpleOrder.eq_bot_or_eq_top (Subgroup.zpowers (g ^ n))).symm.imp ?_ ?_ · intro h have hgo := orderOf_pow (n := n) g rw [orderOf_eq_card_of_forall_mem_zpowers hg, Nat.gcd_eq_right_iff_dvd.1 hn, orderOf_eq_card_of_forall_mem_zpowers, eq_comm, Nat.div_eq_iff_eq_mul_left (Nat.pos_of_dvd_of_pos hn h0) hn] at hgo - · exact (mul_left_cancel₀ (ne_of_gt h0) ((mul_one (Fintype.card α)).trans hgo)).symm + · exact (mul_left_cancel₀ (ne_of_gt h0) ((mul_one (Nat.card α)).trans hgo)).symm · intro x rw [h] exact Subgroup.mem_top _ @@ -590,13 +575,13 @@ end CommGroup end IsSimpleGroup @[to_additive] -theorem CommGroup.is_simple_iff_isCyclic_and_prime_card [Fintype α] [CommGroup α] : - IsSimpleGroup α ↔ IsCyclic α ∧ (Fintype.card α).Prime := by +theorem CommGroup.is_simple_iff_isCyclic_and_prime_card [Finite α] [CommGroup α] : + IsSimpleGroup α ↔ IsCyclic α ∧ (Nat.card α).Prime := by constructor · intro h exact ⟨IsSimpleGroup.isCyclic, IsSimpleGroup.prime_card⟩ · rintro ⟨_, hp⟩ - haveI : Fact (Fintype.card α).Prime := ⟨hp⟩ + haveI : Fact (Nat.card α).Prime := ⟨hp⟩ exact isSimpleGroup_of_prime_card rfl section SpecificInstances @@ -617,24 +602,23 @@ section Exponent open Monoid @[to_additive] -theorem IsCyclic.exponent_eq_card [Group α] [IsCyclic α] [Fintype α] : - exponent α = Fintype.card α := by +theorem IsCyclic.exponent_eq_card [Group α] [IsCyclic α] : + exponent α = Nat.card α := by obtain ⟨g, hg⟩ := IsCyclic.exists_generator (α := α) - apply Nat.dvd_antisymm - · rw [← lcm_orderOf_eq_exponent, Finset.lcm_dvd_iff] - exact fun b _ => orderOf_dvd_card + apply Nat.dvd_antisymm Group.exponent_dvd_nat_card rw [← orderOf_eq_card_of_forall_mem_zpowers hg] exact order_dvd_exponent _ @[to_additive] -theorem IsCyclic.of_exponent_eq_card [CommGroup α] [Fintype α] (h : exponent α = Fintype.card α) : +theorem IsCyclic.of_exponent_eq_card [CommGroup α] [Finite α] (h : exponent α = Nat.card α) : IsCyclic α := + let ⟨_⟩ := nonempty_fintype α let ⟨g, _, hg⟩ := Finset.mem_image.mp (Finset.max'_mem _ _) isCyclic_of_orderOf_eq_card g <| hg.trans <| exponent_eq_max'_orderOf.symm.trans h @[to_additive] -theorem IsCyclic.iff_exponent_eq_card [CommGroup α] [Fintype α] : - IsCyclic α ↔ exponent α = Fintype.card α := +theorem IsCyclic.iff_exponent_eq_card [CommGroup α] [Finite α] : + IsCyclic α ↔ exponent α = Nat.card α := ⟨fun _ => IsCyclic.exponent_eq_card, IsCyclic.of_exponent_eq_card⟩ @[to_additive] @@ -645,29 +629,26 @@ theorem IsCyclic.exponent_eq_zero_of_infinite [Group α] [IsCyclic α] [Infinite @[simp] protected theorem ZMod.exponent (n : ℕ) : AddMonoid.exponent (ZMod n) = n := by - cases n - · rw [IsAddCyclic.exponent_eq_zero_of_infinite] - · rw [IsAddCyclic.exponent_eq_card, card] + rw [IsAddCyclic.exponent_eq_card, Nat.card_zmod] /-- A group of order `p ^ 2` is not cyclic if and only if its exponent is `p`. -/ @[to_additive] lemma not_isCyclic_iff_exponent_eq_prime [Group α] {p : ℕ} (hp : p.Prime) (hα : Nat.card α = p ^ 2) : ¬ IsCyclic α ↔ Monoid.exponent α = p := by -- G is a nontrivial fintype of cardinality `p ^ 2` - 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 + have : Finite α := Nat.finite_of_card_ne_zero (hα ▸ pow_ne_zero 2 hp.ne_zero) + have : Nontrivial α := Finite.one_lt_card_iff_nontrivial.mp + (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 - obtain (rfl|rfl) := eq_zero_or_one_of_sq_eq_self <| hα' ▸ h_exp ▸ (h_cyc.exponent_eq_card).symm + obtain (rfl|rfl) := eq_zero_or_one_of_sq_eq_self <| hα ▸ h_exp ▸ (h_cyc.exponent_eq_card).symm · exact Nat.not_prime_zero hp · exact Nat.not_prime_one hp⟩ /- we must show every non-identity element has order `p`. By Lagrange's theorem, the only possible orders of `g` are `1`, `p`, or `p ^ 2`. It can't be the former because `g ≠ 1`, and it can't the latter because the group isn't cyclic. -/ - have := (Nat.mem_divisors (m := p ^ 2)).mpr ⟨hα' ▸ orderOf_dvd_card (x := g), by aesop⟩ + have := (Nat.mem_divisors (m := p ^ 2)).mpr ⟨hα ▸ orderOf_dvd_natCard (x := g), by aesop⟩ simp? [Nat.divisors_prime_pow hp 2] at this says simp only [Nat.divisors_prime_pow hp 2, Nat.reduceAdd, Finset.mem_map, Finset.mem_range, Function.Embedding.coeFn_mk] at this @@ -729,12 +710,11 @@ noncomputable def mulEquivOfCyclicCardEq [Group G] [Group G'] [hG : IsCyclic G] /-- Two groups of the same prime cardinality are isomorphic. -/ @[to_additive "Two additive groups of the same prime cardinality are isomorphic."] -noncomputable def mulEquivOfPrimeCardEq {p : ℕ} [Fintype G] [Fintype G'] [Group G] [Group G'] - [Fact p.Prime] (hG : Fintype.card G = p) (hH : Fintype.card G' = p) : G ≃* G' := by +noncomputable def mulEquivOfPrimeCardEq {p : ℕ} [Group G] [Group G'] + [Fact p.Prime] (hG : Nat.card G = p) (hH : Nat.card G' = p) : G ≃* G' := by have hGcyc := isCyclic_of_prime_card hG have hHcyc := isCyclic_of_prime_card hH apply mulEquivOfCyclicCardEq - rw [← Nat.card_eq_fintype_card] at hG hH exact hG.trans hH.symm end ZMod diff --git a/Mathlib/GroupTheory/SpecificGroups/Quaternion.lean b/Mathlib/GroupTheory/SpecificGroups/Quaternion.lean index 0db9e8be9fe21..3ce9f9ecba4d9 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Quaternion.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Quaternion.lean @@ -210,7 +210,7 @@ theorem orderOf_xa [NeZero n] (i : ZMod (2 * n)) : orderOf (xa i) = 4 := by /-- In the special case `n = 1`, `Quaternion 1` is a cyclic group (of order `4`). -/ theorem quaternionGroup_one_isCyclic : IsCyclic (QuaternionGroup 1) := by apply isCyclic_of_orderOf_eq_card - · rw [card, mul_one] + · rw [Nat.card_eq_fintype_card, card, mul_one] exact orderOf_xa 0 /-- If `0 < n`, then `a 1` has order `2 * n`. diff --git a/Mathlib/NumberTheory/Cyclotomic/CyclotomicCharacter.lean b/Mathlib/NumberTheory/Cyclotomic/CyclotomicCharacter.lean index 6b59931b1a4c7..dde29ae2b1750 100644 --- a/Mathlib/NumberTheory/Cyclotomic/CyclotomicCharacter.lean +++ b/Mathlib/NumberTheory/Cyclotomic/CyclotomicCharacter.lean @@ -119,7 +119,7 @@ theorem toFun_spec'' (g : L ≃+* L) {n : ℕ} [NeZero n] {t : L} (ht : IsPrimit /-- If g(t)=t^c for all roots of unity, then c=χ(g). -/ theorem toFun_unique (g : L ≃+* L) (c : ZMod (Fintype.card (rootsOfUnity n L))) (hc : ∀ t : rootsOfUnity n L, g (t : Lˣ) = (t ^ c.val : Lˣ)) : c = χ₀ n g := by - apply IsCyclic.ext rfl (fun ζ ↦ ?_) + apply IsCyclic.ext Nat.card_eq_fintype_card (fun ζ ↦ ?_) specialize hc ζ suffices ((ζ ^ c.val : Lˣ) : L) = (ζ ^ (χ₀ n g).val : Lˣ) by exact_mod_cast this rw [← toFun_spec g ζ, hc] diff --git a/Mathlib/NumberTheory/LucasPrimality.lean b/Mathlib/NumberTheory/LucasPrimality.lean index 47887d769b36b..34deef6e5fc10 100644 --- a/Mathlib/NumberTheory/LucasPrimality.lean +++ b/Mathlib/NumberTheory/LucasPrimality.lean @@ -54,7 +54,8 @@ theorem reverse_lucas_primality (p : ℕ) (hP : p.Prime) : have : Fact p.Prime := ⟨hP⟩ obtain ⟨g, hg⟩ := IsCyclic.exists_generator (α := (ZMod p)ˣ) have h1 : orderOf g = p - 1 := by - rwa [orderOf_eq_card_of_forall_mem_zpowers hg, ← Nat.prime_iff_card_units] + rwa [orderOf_eq_card_of_forall_mem_zpowers hg, Nat.card_eq_fintype_card, + ← Nat.prime_iff_card_units] have h2 := tsub_pos_iff_lt.2 hP.one_lt rw [← orderOf_injective (Units.coeHom _) Units.ext _, orderOf_eq_iff h2] at h1 refine ⟨g, h1.1, fun q hq hqd ↦ ?_⟩ diff --git a/Mathlib/NumberTheory/MulChar/Lemmas.lean b/Mathlib/NumberTheory/MulChar/Lemmas.lean index 205530c0d8bfa..1a90d56bec149 100644 --- a/Mathlib/NumberTheory/MulChar/Lemmas.lean +++ b/Mathlib/NumberTheory/MulChar/Lemmas.lean @@ -82,7 +82,7 @@ noncomputable def ofRootOfUnity {ζ : Rˣ} (hζ : ζ ∈ rootsOfUnity (Fintype.c have : orderOf ζ ∣ Fintype.card Mˣ := orderOf_dvd_iff_pow_eq_one.mpr <| (mem_rootsOfUnity _ ζ).mp hζ refine ofUnitHom <| monoidHomOfForallMemZpowers hg <| this.trans <| dvd_of_eq ?_ - rw [orderOf_generator_eq_natCard hg, Nat.card_eq_fintype_card] + rw [orderOf_eq_card_of_forall_mem_zpowers hg, Nat.card_eq_fintype_card] lemma ofRootOfUnity_spec {ζ : Rˣ} (hζ : ζ ∈ rootsOfUnity (Fintype.card Mˣ) R) {g : Mˣ} (hg : ∀ x, x ∈ Subgroup.zpowers g) : diff --git a/Mathlib/RingTheory/RootsOfUnity/Basic.lean b/Mathlib/RingTheory/RootsOfUnity/Basic.lean index 6edff0a269a73..a4a2e5f1d6ed4 100644 --- a/Mathlib/RingTheory/RootsOfUnity/Basic.lean +++ b/Mathlib/RingTheory/RootsOfUnity/Basic.lean @@ -251,12 +251,12 @@ namespace IsCyclic `n` into another group `G'` to the group of `n`th roots of unity in `G'` determined by a generator `g` of `G`. It sends `φ : G →* G'` to `φ g`. -/ noncomputable -def monoidHomMulEquivRootsOfUnityOfGenerator {G : Type*} [CommGroup G] [Fintype G] {g : G} +def monoidHomMulEquivRootsOfUnityOfGenerator {G : Type*} [CommGroup G] {g : G} (hg : ∀ (x : G), x ∈ Subgroup.zpowers g) (G' : Type*) [CommGroup G'] : - (G →* G') ≃* rootsOfUnity (Fintype.card G) G' where + (G →* G') ≃* rootsOfUnity (Nat.card G) G' where toFun φ := ⟨(IsUnit.map φ <| Group.isUnit g).unit, by simp only [mem_rootsOfUnity, Units.ext_iff, Units.val_pow_eq_pow_val, IsUnit.unit_spec, - ← map_pow, pow_card_eq_one, map_one, Units.val_one]⟩ + ← map_pow, pow_card_eq_one', map_one, Units.val_one]⟩ invFun ζ := monoidHomOfForallMemZpowers hg (g' := (ζ.val : G')) <| by simpa only [orderOf_eq_card_of_forall_mem_zpowers hg, orderOf_dvd_iff_pow_eq_one, ← Units.val_pow_eq_pow_val, Units.val_eq_one] using ζ.prop @@ -270,12 +270,11 @@ def monoidHomMulEquivRootsOfUnityOfGenerator {G : Type*} [CommGroup G] [Fintype /-- The group of group homomorphisms from a finite cyclic group `G` of order `n` into another group `G'` is (noncanonically) isomorphic to the group of `n`th roots of unity in `G'`. -/ -lemma monoidHom_mulEquiv_rootsOfUnity (G : Type*) [CommGroup G] [Finite G] [IsCyclic G] +lemma monoidHom_mulEquiv_rootsOfUnity (G : Type*) [CommGroup G] [IsCyclic G] (G' : Type*) [CommGroup G'] : Nonempty <| (G →* G') ≃* rootsOfUnity (Nat.card G) G' := by obtain ⟨g, hg⟩ := IsCyclic.exists_generator (α := G) - have : Fintype G := Fintype.ofFinite _ - exact ⟨Nat.card_eq_fintype_card (α := G) ▸ monoidHomMulEquivRootsOfUnityOfGenerator hg G'⟩ + exact ⟨monoidHomMulEquivRootsOfUnityOfGenerator hg G'⟩ end IsCyclic diff --git a/Mathlib/RingTheory/RootsOfUnity/EnoughRootsOfUnity.lean b/Mathlib/RingTheory/RootsOfUnity/EnoughRootsOfUnity.lean index 6782bf665be6b..df9beb8a7e6db 100644 --- a/Mathlib/RingTheory/RootsOfUnity/EnoughRootsOfUnity.lean +++ b/Mathlib/RingTheory/RootsOfUnity/EnoughRootsOfUnity.lean @@ -64,8 +64,7 @@ lemma natCard_rootsOfUnity (M : Type*) [CommMonoid M] (n : ℕ) [NeZero n] [HasEnoughRootsOfUnity M n] : Nat.card (rootsOfUnity n M) = n := by obtain ⟨ζ, h⟩ := exists_primitiveRoot M n - have : Fintype <| rootsOfUnity n M := Fintype.ofFinite _ - rw [Nat.card_eq_fintype_card, ← IsCyclic.exponent_eq_card] + rw [← IsCyclic.exponent_eq_card] refine dvd_antisymm ?_ ?_ · exact Monoid.exponent_dvd_of_forall_pow_eq_one fun g ↦ OneMemClass.coe_eq_one.mp g.prop · nth_rewrite 1 [h.eq_orderOf] From 0995fe680d3970679d4334b0e993235073f72511 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 18 Nov 2024 10:08:47 +0000 Subject: [PATCH 053/829] chore: missing deprecations from #18967 (#19081) --- Mathlib/Algebra/Order/Ring/Basic.lean | 28 +++++++++---------- .../Analysis/Distribution/SchwartzSpace.lean | 2 +- .../Trigonometric/Bounds.lean | 2 +- .../MeasureTheory/Integral/PeakFunction.lean | 4 +-- 4 files changed, 18 insertions(+), 18 deletions(-) diff --git a/Mathlib/Algebra/Order/Ring/Basic.lean b/Mathlib/Algebra/Order/Ring/Basic.lean index 542142b07165f..65844e24d7f30 100644 --- a/Mathlib/Algebra/Order/Ring/Basic.lean +++ b/Mathlib/Algebra/Order/Ring/Basic.lean @@ -67,7 +67,7 @@ attribute [bound] pow_le_one₀ one_le_pow₀ @[deprecated (since := "2024-10-04")] alias pow_le_pow_right := pow_le_pow_right₀ @[deprecated (since := "2024-10-04")] alias le_self_pow := le_self_pow₀ -@[mono, gcongr, bound] +@[deprecated pow_le_pow_left₀ (since := "2024-11-13")] theorem pow_le_pow_left {a b : R} (ha : 0 ≤ a) (hab : a ≤ b) : ∀ n, a ^ n ≤ b ^ n := pow_le_pow_left₀ ha hab @@ -143,49 +143,49 @@ end StrictOrderedRing section LinearOrderedSemiring variable [LinearOrderedSemiring R] {a b : R} {m n : ℕ} -@[deprecated (since := "2024-11-12")] +@[deprecated pow_le_pow_iff_left₀ (since := "2024-11-12")] lemma pow_le_pow_iff_left (ha : 0 ≤ a) (hb : 0 ≤ b) (hn : n ≠ 0) : a ^ n ≤ b ^ n ↔ a ≤ b := pow_le_pow_iff_left₀ ha hb hn -@[deprecated (since := "2024-11-12")] +@[deprecated pow_lt_pow_iff_left₀ (since := "2024-11-12")] lemma pow_lt_pow_iff_left (ha : 0 ≤ a) (hb : 0 ≤ b) (hn : n ≠ 0) : a ^ n < b ^ n ↔ a < b := pow_lt_pow_iff_left₀ ha hb hn -@[deprecated (since := "2024-11-12")] +@[deprecated pow_left_inj₀ (since := "2024-11-12")] lemma pow_left_inj (ha : 0 ≤ a) (hb : 0 ≤ b) (hn : n ≠ 0) : a ^ n = b ^ n ↔ a = b := pow_left_inj₀ ha hb hn -@[deprecated (since := "2024-11-12")] +@[deprecated pow_right_injective₀ (since := "2024-11-12")] lemma pow_right_injective (ha₀ : 0 < a) (ha₁ : a ≠ 1) : Injective (a ^ ·) := pow_right_injective₀ ha₀ ha₁ -@[deprecated (since := "2024-11-12")] +@[deprecated pow_right_inj₀ (since := "2024-11-12")] lemma pow_right_inj (ha₀ : 0 < a) (ha₁ : a ≠ 1) : a ^ m = a ^ n ↔ m = n := pow_right_inj₀ ha₀ ha₁ -@[deprecated (since := "2024-11-12")] +@[deprecated sq_le_one_iff₀ (since := "2024-11-12")] theorem sq_le_one_iff {a : R} (ha : 0 ≤ a) : a ^ 2 ≤ 1 ↔ a ≤ 1 := sq_le_one_iff₀ ha -@[deprecated (since := "2024-11-12")] +@[deprecated sq_lt_one_iff₀ (since := "2024-11-12")] theorem sq_lt_one_iff {a : R} (ha : 0 ≤ a) : a ^ 2 < 1 ↔ a < 1 := sq_lt_one_iff₀ ha -@[deprecated (since := "2024-11-12")] +@[deprecated one_le_sq_iff₀ (since := "2024-11-12")] theorem one_le_sq_iff {a : R} (ha : 0 ≤ a) : 1 ≤ a ^ 2 ↔ 1 ≤ a := one_le_sq_iff₀ ha -@[deprecated (since := "2024-11-12")] +@[deprecated one_lt_sq_iff₀ (since := "2024-11-12")] theorem one_lt_sq_iff {a : R} (ha : 0 ≤ a) : 1 < a ^ 2 ↔ 1 < a := one_lt_sq_iff₀ ha -@[deprecated (since := "2024-11-12")] +@[deprecated lt_of_pow_lt_pow_left₀ (since := "2024-11-12")] theorem lt_of_pow_lt_pow_left (n : ℕ) (hb : 0 ≤ b) (h : a ^ n < b ^ n) : a < b := lt_of_pow_lt_pow_left₀ n hb h -@[deprecated (since := "2024-11-12")] +@[deprecated le_of_pow_le_pow_left₀ (since := "2024-11-12")] theorem le_of_pow_le_pow_left (hn : n ≠ 0) (hb : 0 ≤ b) (h : a ^ n ≤ b ^ n) : a ≤ b := le_of_pow_le_pow_left₀ hn hb h -@[deprecated (since := "2024-11-12")] +@[deprecated sq_eq_sq₀ (since := "2024-11-12")] theorem sq_eq_sq {a b : R} (ha : 0 ≤ a) (hb : 0 ≤ b) : a ^ 2 = b ^ 2 ↔ a = b := sq_eq_sq₀ ha hb -@[deprecated (since := "2024-11-12")] +@[deprecated lt_of_mul_self_lt_mul_self₀ (since := "2024-11-12")] theorem lt_of_mul_self_lt_mul_self (hb : 0 ≤ b) : a * a < b * b → a < b := lt_of_mul_self_lt_mul_self₀ hb diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index f4d8fedbb56ec..6debdf4d51000 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -869,7 +869,7 @@ def compCLM {g : D → E} (hg : g.HasTemperateGrowth) 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 := - pow_le_pow_left (norm_nonneg _) (by simp only [zero_le_one, le_add_iff_nonneg_left]) _ + pow_le_pow_left₀ (norm_nonneg _) (by simp only [zero_le_one, le_add_iff_nonneg_left]) _ refine le_trans (mul_le_mul hxk this (by positivity) (by positivity)) ?_ have rearrange : (1 + ‖x‖) ^ k * diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Bounds.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Bounds.lean index bddce0ffb926a..9dd246c8eca0e 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Bounds.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Bounds.lean @@ -135,7 +135,7 @@ lemma cos_le_one_sub_mul_cos_sq (hx : |x| ≤ π) : cos x ≤ 1 - 2 / π ^ 2 * x case inr => simpa using this (by rwa [abs_neg]) <| neg_nonneg.2 <| le_of_not_le hx₀ rw [abs_of_nonneg hx₀] at hx have : x / π ≤ sin (x / 2) := by simpa using mul_le_sin (x := x / 2) (by positivity) (by linarith) - have := (pow_le_pow_left (by positivity) this 2).trans_eq (sin_sq_eq_half_sub _) + have := (pow_le_pow_left₀ (by positivity) this 2).trans_eq (sin_sq_eq_half_sub _) ring_nf at this ⊢ linarith diff --git a/Mathlib/MeasureTheory/Integral/PeakFunction.lean b/Mathlib/MeasureTheory/Integral/PeakFunction.lean index e9adaff6ffe0a..7d17886a1c7d2 100644 --- a/Mathlib/MeasureTheory/Integral/PeakFunction.lean +++ b/Mathlib/MeasureTheory/Integral/PeakFunction.lean @@ -317,12 +317,12 @@ theorem tendsto_setIntegral_pow_smul_of_unique_maximum_of_isCompact_of_measure_n exact lt_of_le_of_lt (measure_mono inter_subset_right) hs.measure_lt_top · exact (I n).mono inter_subset_right le_rfl · intro x hx - exact pow_le_pow_left t'_pos.le (le_of_lt (hv hx)) _ + exact pow_le_pow_left₀ t'_pos.le (hv hx).le _ _ ≤ ∫ y in s, c y ^ n ∂μ := setIntegral_mono_set (I n) (J n) (Eventually.of_forall inter_subset_right) simp_rw [φ, ← div_eq_inv_mul, div_pow, div_div] apply div_le_div (pow_nonneg t_pos n) _ _ B - · exact pow_le_pow_left (hnc _ hx.1) (ht x hx) _ + · exact pow_le_pow_left₀ (hnc _ hx.1) (ht x hx) _ · apply mul_pos (pow_pos (t_pos.trans_lt tt') _) (ENNReal.toReal_pos (hμ v v_open x₀_v).ne' _) have : μ (v ∩ s) ≤ μ s := measure_mono inter_subset_right exact ne_of_lt (lt_of_le_of_lt this hs.measure_lt_top) From 15fca27644d5971140ae06ec4f2a42a6579668ca Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Mon, 18 Nov 2024 10:08:49 +0000 Subject: [PATCH 054/829] refactor(GroupTheory/Index): Remove unnecessary coercions to `Set` in `index_ker` and `relindex_ker` (#19105) This PR removes unnecessary coercions to `Set` in `index_ker` and `relindex_ker`. --- Mathlib/GroupTheory/Index.lean | 7 ++----- 1 file changed, 2 insertions(+), 5 deletions(-) diff --git a/Mathlib/GroupTheory/Index.lean b/Mathlib/GroupTheory/Index.lean index 0eaa15ed8fab6..8357feb56afc9 100644 --- a/Mathlib/GroupTheory/Index.lean +++ b/Mathlib/GroupTheory/Index.lean @@ -208,15 +208,12 @@ theorem relindex_bot_right : H.relindex ⊥ = 1 := by rw [relindex, subgroupOf_b theorem relindex_self : H.relindex H = 1 := by rw [relindex, subgroupOf_self, index_top] @[to_additive] -theorem index_ker (f : G →* G') : f.ker.index = Nat.card (Set.range f) := by +theorem index_ker (f : G →* G') : f.ker.index = Nat.card f.range := by rw [← MonoidHom.comap_bot, index_comap, relindex_bot_left] - rfl @[to_additive] -theorem relindex_ker (f : G →* G') (K : Subgroup G) : - f.ker.relindex K = Nat.card (f '' K) := by +theorem relindex_ker (f : G →* G') : f.ker.relindex K = Nat.card (K.map f) := by rw [← MonoidHom.comap_bot, relindex_comap, relindex_bot_left] - rfl @[to_additive (attr := simp) card_mul_index] theorem card_mul_index : Nat.card H * H.index = Nat.card G := by From cdbc43eebed0f99498b560c972b329f6fececb61 Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 18 Nov 2024 10:08:50 +0000 Subject: [PATCH 055/829] feat(stacks/kerodon): automatically update doc-string with link (#19131) [Zulip discussion](https://leanprover.zulipchat.com/#narrow/channel/113488-general/topic/stacks.20tags/near/482646447) --- Mathlib/Tactic/StacksAttribute.lean | 19 +++++++++++++------ 1 file changed, 13 insertions(+), 6 deletions(-) diff --git a/Mathlib/Tactic/StacksAttribute.lean b/Mathlib/Tactic/StacksAttribute.lean index f3c775c5ce97a..c8f7d829151e7 100644 --- a/Mathlib/Tactic/StacksAttribute.lean +++ b/Mathlib/Tactic/StacksAttribute.lean @@ -143,12 +143,19 @@ syntax (name := stacksTag) stacksTagDB stacksTagParser (ppSpace str)? : attr initialize Lean.registerBuiltinAttribute { name := `stacksTag descr := "Apply a Stacks or Kerodon project tag to a theorem." - add := fun decl stx _attrKind => match stx with - | `(attr| stacks $tag $[$comment]?) => do - addTagEntry decl .stacks (← tag.getStacksTag) <| (comment.map (·.getString)).getD "" - | `(attr| kerodon $tag $[$comment]?) => do - addTagEntry decl .kerodon (← tag.getStacksTag) <| (comment.map (·.getString)).getD "" - | _ => throwUnsupportedSyntax + add := fun decl stx _attrKind => do + let oldDoc := (← findDocString? (← getEnv) decl).getD "" + let (SorK, database, url, tag, comment) := ← match stx with + | `(attr| stacks $tag $[$comment]?) => + return ("Stacks", Database.stacks, "https://stacks.math.columbia.edu/tag", tag, comment) + | `(attr| kerodon $tag $[$comment]?) => + return ("Kerodon", Database.kerodon, "https://kerodon.net/tag", tag, comment) + | _ => throwUnsupportedSyntax + let tagStr ← tag.getStacksTag + let comment := (comment.map (·.getString)).getD "" + let newDoc := [oldDoc, s!"[{SorK} Tag {tagStr}]({url}/{tagStr})", comment] + addDocString decl <| "\n\n".intercalate (newDoc.filter (· != "")) + addTagEntry decl database tagStr <| comment } end Mathlib.StacksTag From dcaf73cb43b061ca74fe8eb480fc0a4376dcb43e Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Mon, 18 Nov 2024 10:53:35 +0000 Subject: [PATCH 056/829] feat: add some lemmas about Monoid.exponent and Int-powers (#19079) --- Mathlib/GroupTheory/Exponent.lean | 12 +++++++++++- Mathlib/GroupTheory/OrderOfElement.lean | 4 ++++ 2 files changed, 15 insertions(+), 1 deletion(-) diff --git a/Mathlib/GroupTheory/Exponent.lean b/Mathlib/GroupTheory/Exponent.lean index de0f7edebeedc..6ee49e4b8a6ac 100644 --- a/Mathlib/GroupTheory/Exponent.lean +++ b/Mathlib/GroupTheory/Exponent.lean @@ -494,7 +494,7 @@ end Monoid section Group -variable [Group G] +variable [Group G] {n m : ℤ} @[to_additive (attr := deprecated Monoid.one_lt_exponent (since := "2024-02-17")) AddGroup.one_lt_exponent] @@ -522,6 +522,16 @@ theorem Subgroup.exponent_top : Monoid.exponent (⊤ : Subgroup G) = Monoid.expo theorem Subgroup.pow_exponent_eq_one {H : Subgroup G} {g : G} (g_in_H : g ∈ H) : g ^ Monoid.exponent H = 1 := exponent_toSubmonoid H ▸ Submonoid.pow_exponent_eq_one g_in_H +@[to_additive] +theorem Group.exponent_dvd_iff_forall_zpow_eq_one : + (Monoid.exponent G : ℤ) ∣ n ↔ ∀ g : G, g ^ n = 1 := by + simp_rw [Int.natCast_dvd, Monoid.exponent_dvd_iff_forall_pow_eq_one, pow_natAbs_eq_one] + +@[to_additive] +theorem Group.exponent_dvd_sub_iff_zpow_eq_zpow : + (Monoid.exponent G : ℤ) ∣ n - m ↔ ∀ g : G, g ^ n = g ^ m := by + simp_rw [Group.exponent_dvd_iff_forall_zpow_eq_one, zpow_sub, mul_inv_eq_one] + end Group section PiProd diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index 89087b4c7e581..8142e71f6f653 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -572,6 +572,10 @@ theorem orderOf_dvd_iff_zpow_eq_one : (orderOf x : ℤ) ∣ i ↔ x ^ i = 1 := b @[to_additive (attr := simp)] theorem orderOf_inv (x : G) : orderOf x⁻¹ = orderOf x := by simp [orderOf_eq_orderOf_iff] +@[to_additive] +theorem orderOf_dvd_sub_iff_zpow_eq_zpow {a b : ℤ} : (orderOf x : ℤ) ∣ a - b ↔ x ^ a = x ^ b := by + rw [orderOf_dvd_iff_zpow_eq_one, zpow_sub, mul_inv_eq_one] + namespace Subgroup variable {H : Subgroup G} From 9b575f298aabca92e06ebc272b867dfd8c0e17b5 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Mon, 18 Nov 2024 10:53:37 +0000 Subject: [PATCH 057/829] =?UTF-8?q?feat:=20`NormedLatticeAddCommGroup`=20i?= =?UTF-8?q?nstances=20for=20`=E2=84=A4`=20and=20`=E2=84=9A`=20(#19143)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/Analysis/Normed/Order/Lattice.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/Mathlib/Analysis/Normed/Order/Lattice.lean b/Mathlib/Analysis/Normed/Order/Lattice.lean index de21daec89f71..ce40831eba4c7 100644 --- a/Mathlib/Analysis/Normed/Order/Lattice.lean +++ b/Mathlib/Analysis/Normed/Order/Lattice.lean @@ -73,6 +73,14 @@ class NormedLatticeAddCommGroup (α : Type*) extends NormedAddCommGroup α, Lattice α, HasSolidNorm α where add_le_add_left : ∀ a b : α, a ≤ b → ∀ c : α, c + a ≤ c + b +instance Int.normedLatticeAddCommGroup : NormedLatticeAddCommGroup ℤ where + solid x y h := by simpa [← Int.norm_cast_real, ← Int.cast_abs] using h + add_le_add_left _ _ := add_le_add_left + +instance Rat.normedLatticeAddCommGroup : NormedLatticeAddCommGroup ℚ where + solid x y h := by simpa [← Rat.norm_cast_real, ← Rat.cast_abs] using h + add_le_add_left _ _ := add_le_add_left + instance Real.normedLatticeAddCommGroup : NormedLatticeAddCommGroup ℝ where add_le_add_left _ _ h _ := add_le_add le_rfl h From 82e5e81af635e2f3ec506588eff545b0ccf5751d Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Mon, 18 Nov 2024 10:53:38 +0000 Subject: [PATCH 058/829] refactor(NumberTheory/LSeries): merge `QuadraticNonvanishing` into `Nonvanishing` (#19185) The proof of non-vanishing of Dirichlet L-functions on `re s = 1` was written in two chunks, since the case of quadratic characters at `s = 1` needs to be handled separately. This PR merges together the two files, and makes one more theorem private since it is superseded by more general results proved later on, rendering the case split in the proof invisible to end-users. --- Mathlib.lean | 1 - .../NumberTheory/LSeries/Nonvanishing.lean | 218 ++++++++++++++++-- .../LSeries/QuadraticNonvanishing.lean | 186 --------------- 3 files changed, 198 insertions(+), 207 deletions(-) delete mode 100644 Mathlib/NumberTheory/LSeries/QuadraticNonvanishing.lean diff --git a/Mathlib.lean b/Mathlib.lean index 9dbb62dd6373b..e55b078c41c8f 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3754,7 +3754,6 @@ import Mathlib.NumberTheory.LSeries.Linearity import Mathlib.NumberTheory.LSeries.MellinEqDirichlet import Mathlib.NumberTheory.LSeries.Nonvanishing import Mathlib.NumberTheory.LSeries.Positivity -import Mathlib.NumberTheory.LSeries.QuadraticNonvanishing import Mathlib.NumberTheory.LSeries.RiemannZeta import Mathlib.NumberTheory.LSeries.ZMod import Mathlib.NumberTheory.LegendreSymbol.AddCharacter diff --git a/Mathlib/NumberTheory/LSeries/Nonvanishing.lean b/Mathlib/NumberTheory/LSeries/Nonvanishing.lean index f3b7eb0bc0832..7ed3e3b3b2c37 100644 --- a/Mathlib/NumberTheory/LSeries/Nonvanishing.lean +++ b/Mathlib/NumberTheory/LSeries/Nonvanishing.lean @@ -1,33 +1,219 @@ /- Copyright (c) 2024 Michael Stoll. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Michael Stoll +Authors: Michael Stoll, David Loeffler -/ import Mathlib.Analysis.SpecialFunctions.Complex.LogBounds import Mathlib.NumberTheory.Harmonic.ZetaAsymp -import Mathlib.NumberTheory.LSeries.QuadraticNonvanishing +import Mathlib.NumberTheory.LSeries.Dirichlet +import Mathlib.NumberTheory.LSeries.DirichletContinuation +import Mathlib.NumberTheory.LSeries.Positivity /-! # The L-function of a Dirichlet character does not vanish on Re(s) ≥ 1 The main result in this file is `DirichletCharacter.Lfunction_ne_zero_of_one_le_re`: -if `χ` is a Dirichlet character and `s.re ≥ 1` and either `χ` is nontrivial or `s ≠ 1`, +if `χ` is a Dirichlet character, `s ∈ ℂ` with `1 ≤ s.re`, and either `χ` is nontrivial or `s ≠ 1`, then the L-function of `χ` does not vanish at `s`. As a consequence, we have the corresponding statement for the Riemann ζ function: -`riemannZeta_ne_zero_of_one_le_re`. +`riemannZeta_ne_zero_of_one_le_re` (which does not require `s ≠ 1`, since the junk value at `s = 1` +happens to be non-zero). These results are prerequisites for the **Prime Number Theorem** and **Dirichlet's Theorem** on primes in arithmetic progressions. + +## Outline of proofs + +We split into two cases: first, the special case of (non-trivial) quadratic characters at `s = 1`; +then the remaining case when either `s ≠ 1` or `χ ^ 2 ≠ 1`. + +The first case is handled using a positivity argument applied to the series `L χ s * ζ s`: we show +that this function has non-negative Dirichlet coefficients, is strictly positive for `s ≫ 0`, but +vanishes at `s = -2`, so it must have a pole somewhere in between. + +The second case is dealt with using the product +`L(χ^0, 1 + x)^3 L(χ, 1 + x + I * y)^4 L(χ^2, 1 + x + 2 * I * y)`, which +we show has absolute value `≥ 1` for all positive `x` and real `y`; if `L(χ, 1 + I * y) = 0` then +this product would have to tend to 0 as `x → 0`, which is a contradiction. +-/ + +/- NB: Many lemmas (and some defs) in this file are private, since they concern properties of +hypothetical objects which we eventually deduce cannot exist. We have only made public the lemmas +whose hypotheses do not turn out to be contradictory. +-/ + +open Complex Asymptotics Topology Filter +open ArithmeticFunction hiding log + +-- We use the ordering on `ℂ` given by comparing real parts for fixed imaginary part +open scoped ComplexOrder + +variable {N : ℕ} + +namespace DirichletCharacter + +section quadratic + +/-! +### Convolution of a Dirichlet character with ζ + +We define `DirichletCharacter.zetaMul χ` to be the arithmetic function obtained by +taking the product (as arithmetic functions = Dirichlet convolution) of the +arithmetic function `ζ` with `χ`. + +We then show that for a quadratic character `χ`, this arithmetic function is multiplicative +and takes nonnegative real values. -/ +/-- The complex-valued arithmetic function that is the convolution of the constant +function `1` with `χ`. -/ +def zetaMul (χ : DirichletCharacter ℂ N) : ArithmeticFunction ℂ := + .zeta * toArithmeticFunction (χ ·) + +/-- The arithmetic function `zetaMul χ` is multiplicative. -/ +lemma isMultiplicative_zetaMul (χ : DirichletCharacter ℂ N) : χ.zetaMul.IsMultiplicative := + isMultiplicative_zeta.natCast.mul <| isMultiplicative_toArithmeticFunction χ + +lemma LSeriesSummable_zetaMul (χ : DirichletCharacter ℂ N) {s : ℂ} (hs : 1 < s.re) : + LSeriesSummable χ.zetaMul s := by + refine ArithmeticFunction.LSeriesSummable_mul (LSeriesSummable_zeta_iff.mpr hs) <| + LSeriesSummable_of_bounded_of_one_lt_re (m := 1) (fun n hn ↦ ?_) hs + simpa only [toArithmeticFunction, coe_mk, hn, ↓reduceIte, ← Complex.norm_eq_abs] + using norm_le_one χ _ + +lemma zetaMul_prime_pow_nonneg {χ : DirichletCharacter ℂ N} (hχ : χ ^ 2 = 1) {p : ℕ} + (hp : p.Prime) (k : ℕ) : + 0 ≤ zetaMul χ (p ^ k) := by + simp only [zetaMul, toArithmeticFunction, coe_zeta_mul_apply, coe_mk, + Nat.sum_divisors_prime_pow hp, pow_eq_zero_iff', hp.ne_zero, ne_eq, false_and, ↓reduceIte, + Nat.cast_pow, map_pow] + rcases MulChar.isQuadratic_iff_sq_eq_one.mpr hχ p with h | h | h + · refine Finset.sum_nonneg fun i _ ↦ ?_ + simp only [h, le_refl, pow_nonneg] + · refine Finset.sum_nonneg fun i _ ↦ ?_ + simp only [h, one_pow, zero_le_one] + · simp only [h, neg_one_geom_sum] + split_ifs + exacts [le_rfl, zero_le_one] + +/-- `zetaMul χ` takes nonnegative real values when `χ` is a quadratic character. -/ +lemma zetaMul_nonneg {χ : DirichletCharacter ℂ N} (hχ : χ ^ 2 = 1) (n : ℕ) : + 0 ≤ zetaMul χ n := by + rcases eq_or_ne n 0 with rfl | hn + · simp only [ArithmeticFunction.map_zero, le_refl] + · simpa only [χ.isMultiplicative_zetaMul.multiplicative_factorization _ hn] using + Finset.prod_nonneg + fun p hp ↦ zetaMul_prime_pow_nonneg hχ (Nat.prime_of_mem_primeFactors hp) _ + +/- +### "Bad" Dirichlet characters + +Our goal is to show that `L(χ, 1) ≠ 0` when `χ` is a (nontrivial) quadratic Dirichlet character. +To do that, we package the contradictory properties in a (private) structure +`DirichletCharacter.BadChar` and derive further statements eventually leading to a contradiction. + +This entire section is private. +-/ + +/-- The object we're trying to show doesn't exist: A nontrivial quadratic Dirichlet character +whose L-function vanishes at `s = 1`. -/ +private structure BadChar (N : ℕ) [NeZero N] where + /-- The character we want to show cannot exist. -/ + χ : DirichletCharacter ℂ N + χ_ne : χ ≠ 1 + χ_sq : χ ^ 2 = 1 + hχ : χ.LFunction 1 = 0 + +variable [NeZero N] + +namespace BadChar + +/-- The product of the Riemann zeta function with the L-function of `B.χ`. +We will show that `B.F (-2) = 0` but also that `B.F (-2)` must be positive, +giving the desired contradiction. -/ +private noncomputable +def F (B : BadChar N) : ℂ → ℂ := + Function.update (fun s : ℂ ↦ riemannZeta s * LFunction B.χ s) 1 (deriv (LFunction B.χ) 1) + +private lemma F_differentiableAt_of_ne (B : BadChar N) {s : ℂ} (hs : s ≠ 1) : + DifferentiableAt ℂ B.F s := by + apply DifferentiableAt.congr_of_eventuallyEq + · exact (differentiableAt_riemannZeta hs).mul <| differentiableAt_LFunction B.χ s (.inl hs) + · filter_upwards [eventually_ne_nhds hs] with t ht using Function.update_noteq ht .. + +/-- `B.F` agrees with the L-series of `zetaMul χ` on `1 < s.re`. -/ +private lemma F_eq_LSeries (B : BadChar N) {s : ℂ} (hs : 1 < s.re) : + B.F s = LSeries B.χ.zetaMul s := by + rw [F, zetaMul, ← coe_mul, LSeries_convolution'] + · have hs' : s ≠ 1 := fun h ↦ by simp only [h, one_re, lt_self_iff_false] at hs + simp only [ne_eq, hs', not_false_eq_true, Function.update_noteq, B.χ.LFunction_eq_LSeries hs] + congr 1 + · simp_rw [← LSeries_zeta_eq_riemannZeta hs, ← natCoe_apply] + · exact LSeries_congr s B.χ.apply_eq_toArithmeticFunction_apply + -- summability side goals from `LSeries_convolution'` + · exact LSeriesSummable_zeta_iff.mpr hs + · exact (LSeriesSummable_congr _ fun h ↦ (B.χ.apply_eq_toArithmeticFunction_apply h).symm).mpr <| + ZMod.LSeriesSummable_of_one_lt_re B.χ hs + +/-- If `χ` is a bad character, then `F` is an entire function. -/ +private lemma F_differentiable (B : BadChar N) : Differentiable ℂ B.F := by + intro s + rcases ne_or_eq s 1 with hs | rfl + · exact B.F_differentiableAt_of_ne hs + -- now need to deal with `s = 1` + refine (analyticAt_of_differentiable_on_punctured_nhds_of_continuousAt ?_ ?_).differentiableAt + · filter_upwards [self_mem_nhdsWithin] with t ht + exact B.F_differentiableAt_of_ne ht + -- now reduced to showing *continuity* at s = 1 + let G := Function.update (fun s ↦ (s - 1) * riemannZeta s) 1 1 + let H := Function.update (fun s ↦ (B.χ.LFunction s - B.χ.LFunction 1) / (s - 1)) 1 + (deriv B.χ.LFunction 1) + have : B.F = G * H := by + ext1 t + rcases eq_or_ne t 1 with rfl | ht + · simp only [F, G, H, Pi.mul_apply, one_mul, Function.update_same] + · simp only [F, G, H, Function.update_noteq ht, mul_comm _ (riemannZeta _), B.hχ, sub_zero, + Pi.mul_apply, mul_assoc, mul_div_cancel₀ _ (sub_ne_zero.mpr ht)] + rw [this] + apply ContinuousAt.mul + · simpa only [G, continuousAt_update_same] using riemannZeta_residue_one + · exact (B.χ.differentiableAt_LFunction 1 (.inr B.χ_ne)).hasDerivAt.continuousAt_div + +/-- The trivial zero at `s = -2` of the zeta function gives that `F (-2) = 0`. +This is used later to obtain a contradction. -/ +private lemma F_neg_two (B : BadChar N) : B.F (-2 : ℝ) = 0 := by + have := riemannZeta_neg_two_mul_nat_add_one 0 + rw [Nat.cast_zero, zero_add, mul_one] at this + rw [F, ofReal_neg, ofReal_ofNat, Function.update_noteq (mod_cast (by omega : (-2 : ℤ) ≠ 1)), + this, zero_mul] + +end BadChar + +/-- If `χ` is a nontrivial quadratic Dirichlet character, then `L(χ, 1) ≠ 0`. This is private +since it is later superseded by `LFunction_apply_one_ne_zero`. -/ +private theorem LFunction_apply_one_ne_zero_of_quadratic {χ : DirichletCharacter ℂ N} + (hχ : χ ^ 2 = 1) (χ_ne : χ ≠ 1) : + χ.LFunction 1 ≠ 0 := by + intro hL + -- construct a "bad character" and put together a contradiction. + let B : BadChar N := {χ := χ, χ_sq := hχ, hχ := hL, χ_ne := χ_ne} + refine B.F_neg_two.not_gt ?_ + refine ArithmeticFunction.LSeries_positive_of_differentiable_of_eqOn (zetaMul_nonneg hχ) + (χ.isMultiplicative_zetaMul.map_one ▸ zero_lt_one) B.F_differentiable ?_ + (fun _ ↦ B.F_eq_LSeries) _ + exact LSeries.abscissaOfAbsConv_le_of_forall_lt_LSeriesSummable + fun _ a ↦ χ.LSeriesSummable_zetaMul a + +end quadratic + section nonvanishing -open Complex +variable (χ : DirichletCharacter ℂ N) -- This is the key positivity lemma that is used to show that the L-function -- of a Dirichlet character `χ` does not vanish for `s.re ≥ 1` (unless `χ^2 = 1` and `s = 1`). -private lemma re_log_comb_nonneg {a : ℝ} (ha₀ : 0 ≤ a) (ha₁ : a < 1) {z : ℂ} (hz : ‖z‖ = 1) : +private lemma re_log_comb_nonneg' {a : ℝ} (ha₀ : 0 ≤ a) (ha₁ : a < 1) {z : ℂ} (hz : ‖z‖ = 1) : 0 ≤ 3 * (-log (1 - a)).re + 4 * (-log (1 - a * z)).re + (-log (1 - a * z ^ 2)).re := by have hac₀ : ‖(a : ℂ)‖ < 1 := by simp only [norm_eq_abs, abs_ofReal, _root_.abs_of_nonneg ha₀, ha₁] @@ -48,10 +234,6 @@ private lemma re_log_comb_nonneg {a : ℝ} (ha₀ : 0 ≤ a) (ha₁ : a < 1) {z convert (show 0 ≤ 2 * a ^ n * ((z ^ n).re + 1) ^ 2 by positivity) using 1 ring -namespace DirichletCharacter - -variable {N : ℕ} (χ : DirichletCharacter ℂ N) - -- This is the version of the technical positivity lemma for logarithms of Euler factors. private lemma re_log_comb_nonneg {n : ℕ} (hn : 2 ≤ n) {x : ℝ} (hx : 1 < x) (y : ℝ) : 0 ≤ 3 * (-log (1 - (1 : DirichletCharacter ℂ N) n * n ^ (-x : ℂ))).re + @@ -68,7 +250,7 @@ private lemma re_log_comb_nonneg {n : ℕ} (hn : 2 ≤ n) {x : ℝ} (hx : 1 < x) simp only [neg_re, mul_re, I_re, ofReal_re, zero_mul, I_im, ofReal_im, mul_zero, sub_self, neg_zero, Real.rpow_zero, one_mul] rw [MulChar.one_apply hn', one_mul] - convert _root_.re_log_comb_nonneg ha₀ ha₁ hz using 6 + convert re_log_comb_nonneg' ha₀ ha₁ hz using 6 · simp only [ofReal_cpow n.cast_nonneg (-x), ofReal_natCast, ofReal_neg] · congr 2 rw [neg_add, cpow_add _ _ <| mod_cast by omega, ← ofReal_neg, ofReal_cpow n.cast_nonneg (-x), @@ -96,7 +278,7 @@ private lemma one_lt_re_one_add {x : ℝ} (hx : 0 < x) (y : ℝ) : open scoped LSeries.notation in /-- For positive `x` and nonzero `y` and a Dirichlet character `χ` we have that -`|L(χ^0, x)^3 L(χ, x+iy)^4 L(χ^2, x+2iy)| ≥ 1. -/ +`|L(χ^0, 1 + x)^3 L(χ, 1 + x + I * y)^4 L(χ^2, 1 + x + 2 * I * y)| ≥ 1. -/ lemma norm_LSeries_product_ge_one {x : ℝ} (hx : 0 < x) (y : ℝ) : ‖L ↗(1 : DirichletCharacter ℂ N) (1 + x) ^ 3 * L ↗χ (1 + x + I * y) ^ 4 * L ↗(χ ^ 2 :) (1 + x + 2 * I * y)‖ ≥ 1 := by @@ -128,9 +310,6 @@ lemma norm_LFunction_product_ge_one {x : ℝ} (hx : 0 < x) (y : ℝ) : χ.LFunction_eq_LSeries h₁, (χ ^ 2).LFunction_eq_LSeries h₂] exact norm_LSeries_product_ge_one χ hx y -open Asymptotics Topology Filter - -open Homeomorph in lemma LFunctionTrivChar_isBigO_near_one_horizontal : (fun x : ℝ ↦ LFunctionTrivChar N (1 + x)) =O[𝓝[>] 0] fun x ↦ (1 : ℂ) / x := by have : (fun w : ℂ ↦ LFunctionTrivChar N (1 + w)) =O[𝓝[≠] 0] (1 / ·) := by @@ -203,7 +382,7 @@ except when `χ` is trivial and `s = 1` (then `L(χ, s)` has a simple pole at `s theorem LFunction_ne_zero_of_re_eq_one {s : ℂ} (hs : s.re = 1) (hχs : χ ≠ 1 ∨ s ≠ 1) : LFunction χ s ≠ 0 := by by_cases h : χ ^ 2 = 1 ∧ s = 1 - · exact h.2 ▸ LFunction_at_one_ne_zero_of_quadratic h.1 <| hχs.neg_resolve_right h.2 + · exact h.2 ▸ LFunction_apply_one_ne_zero_of_quadratic h.1 <| hχs.neg_resolve_right h.2 · have hs' : s = 1 + I * s.im := by conv_lhs => rw [← re_add_im s, hs, ofReal_one, mul_comm] rw [not_and_or, ← ne_eq, ← ne_eq, hs', add_right_ne_self] at h @@ -224,15 +403,14 @@ variable {χ} in theorem LFunction_apply_one_ne_zero (hχ : χ ≠ 1) : LFunction χ 1 ≠ 0 := LFunction_ne_zero_of_one_le_re χ (.inl hχ) <| one_re ▸ le_rfl -end DirichletCharacter - -open DirichletCharacter in /-- The Riemann Zeta Function does not vanish on the closed half-plane `re s ≥ 1`. (Note that the value at `s = 1` is a junk value, which happens to be nonzero.) -/ -lemma riemannZeta_ne_zero_of_one_le_re ⦃s : ℂ⦄ (hs : 1 ≤ s.re) : +lemma _root_.riemannZeta_ne_zero_of_one_le_re ⦃s : ℂ⦄ (hs : 1 ≤ s.re) : riemannZeta s ≠ 0 := by rcases eq_or_ne s 1 with rfl | hs₀ · exact riemannZeta_one_ne_zero · exact LFunction_modOne_eq (χ := 1) ▸ LFunction_ne_zero_of_one_le_re _ (.inr hs₀) hs end nonvanishing + +end DirichletCharacter diff --git a/Mathlib/NumberTheory/LSeries/QuadraticNonvanishing.lean b/Mathlib/NumberTheory/LSeries/QuadraticNonvanishing.lean deleted file mode 100644 index 77f0a4d9c37ec..0000000000000 --- a/Mathlib/NumberTheory/LSeries/QuadraticNonvanishing.lean +++ /dev/null @@ -1,186 +0,0 @@ -/- -Copyright (c) 2024 Michael Stoll. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: David Loeffler, Michael Stoll --/ -import Mathlib.NumberTheory.LSeries.DirichletContinuation -import Mathlib.NumberTheory.LSeries.Dirichlet -import Mathlib.NumberTheory.LSeries.Positivity - -/-! -# Non-vanishing of `L(χ, 1)` for nontrivial quadratic characters `χ` - -The main result of this file is the statement -`DirichletCharacter.LFunction_at_one_ne_zero_of_quadratic`, which says that if `χ` is -a nontrivial (`χ ≠ 1`) quadratic (`χ^2 = 1`) Dirichlet character, then the value of -its L-function at `s = 1` is nonzero. - -This is an important step in the proof of -*Dirichlet's Theorem on Primes in Arithmetic Progression*. --/ - -namespace DirichletCharacter - -/-! -### Convolution of a Dirichlet character with ζ - -We define `DirichletCharacter.zetaMul χ` to be the arithmetic function obtained by -taking the product (as arithmetic functions = Dirichlet convolution) of the -arithmetic function `ζ` with `χ`. - -We then show that for a quadratic character `χ`, this arithmetic function is multiplicative -and takes nonnegative real values. --/ - -open Complex ArithmeticFunction - -variable {N : ℕ} - -/-- The complex-valued arithmetic function that is the convolution of the constant -function `1` with `χ`. -/ -def zetaMul (χ : DirichletCharacter ℂ N) : ArithmeticFunction ℂ := - .zeta * toArithmeticFunction (χ ·) - -/-- The arithmetic function `zetaMul χ` is multiplicative. -/ -lemma isMultiplicative_zetaMul (χ : DirichletCharacter ℂ N) : χ.zetaMul.IsMultiplicative := - isMultiplicative_zeta.natCast.mul <| isMultiplicative_toArithmeticFunction χ - -lemma LSeriesSummable_zetaMul (χ : DirichletCharacter ℂ N) {s : ℂ} (hs : 1 < s.re) : - LSeriesSummable χ.zetaMul s := by - refine ArithmeticFunction.LSeriesSummable_mul (LSeriesSummable_zeta_iff.mpr hs) <| - LSeriesSummable_of_bounded_of_one_lt_re (m := 1) (fun n hn ↦ ?_) hs - simpa only [toArithmeticFunction, coe_mk, hn, ↓reduceIte, ← Complex.norm_eq_abs] - using norm_le_one χ _ - --- We use the ordering on `ℂ` given by comparing real parts for fixed imaginary part -open scoped ComplexOrder - -lemma zetaMul_prime_pow_nonneg {χ : DirichletCharacter ℂ N} (hχ : χ ^ 2 = 1) {p : ℕ} - (hp : p.Prime) (k : ℕ) : - 0 ≤ zetaMul χ (p ^ k) := by - simp only [zetaMul, toArithmeticFunction, coe_zeta_mul_apply, coe_mk, - Nat.sum_divisors_prime_pow hp, pow_eq_zero_iff', hp.ne_zero, ne_eq, false_and, ↓reduceIte, - Nat.cast_pow, map_pow] - rcases MulChar.isQuadratic_iff_sq_eq_one.mpr hχ p with h | h | h - · refine Finset.sum_nonneg fun i _ ↦ ?_ - simp only [h, le_refl, pow_nonneg] - · refine Finset.sum_nonneg fun i _ ↦ ?_ - simp only [h, one_pow, zero_le_one] - · simp only [h, neg_one_geom_sum] - split_ifs - exacts [le_rfl, zero_le_one] - -/-- `zetaMul χ` takes nonnegative real values when `χ` is a quadratic character. -/ -lemma zetaMul_nonneg {χ : DirichletCharacter ℂ N} (hχ : χ ^ 2 = 1) (n : ℕ) : - 0 ≤ zetaMul χ n := by - rcases eq_or_ne n 0 with rfl | hn - · simp only [ArithmeticFunction.map_zero, le_refl] - · simpa only [χ.isMultiplicative_zetaMul.multiplicative_factorization _ hn] using - Finset.prod_nonneg - fun p hp ↦ zetaMul_prime_pow_nonneg hχ (Nat.prime_of_mem_primeFactors hp) _ - - -/- -### "Bad" Dirichlet characters - -Our goal is to show that `L(χ, 1) ≠ 0` when `χ` is a (nontrivial) quadratic Dirichlet character. -To do that, we package the contradictory properties in a (private) structure -`DirichletCharacter.BadChar` and derive further statements eventually leading to a contradiction. --/ - -/-- The object we're trying to show doesn't exist: A nontrivial quadratic Dirichlet character -whose L-function vanishes at `s = 1`. -/ -private structure BadChar (N : ℕ) [NeZero N] where - /-- The character we want to show cannot exist. -/ - χ : DirichletCharacter ℂ N - χ_ne : χ ≠ 1 - χ_sq : χ ^ 2 = 1 - hχ : χ.LFunction 1 = 0 - -variable {N : ℕ} [NeZero N] - -open Complex DirichletCharacter - -namespace BadChar - -/-- The product of the Riemann zeta function with the L-function of `B.χ`. -We will show that `B.F (-2) = 0` but also that `B.F (-2)` must be positive, -giving the desired contradiction. -/ -private noncomputable -def F (B : BadChar N) : ℂ → ℂ := - Function.update (fun s : ℂ ↦ riemannZeta s * LFunction B.χ s) 1 (deriv (LFunction B.χ) 1) - -private lemma F_differentiableAt_of_ne (B : BadChar N) {s : ℂ} (hs : s ≠ 1) : - DifferentiableAt ℂ B.F s := by - apply DifferentiableAt.congr_of_eventuallyEq - · exact (differentiableAt_riemannZeta hs).mul <| differentiableAt_LFunction B.χ s (.inl hs) - · filter_upwards [eventually_ne_nhds hs] with t ht using Function.update_noteq ht .. - -open ArithmeticFunction in -/-- `B.F` agrees with the L-series of `zetaMul χ` on `1 < s.re`. -/ -private lemma F_eq_LSeries (B : BadChar N) {s : ℂ} (hs : 1 < s.re) : - B.F s = LSeries B.χ.zetaMul s := by - rw [F, zetaMul, ← coe_mul, LSeries_convolution'] - · have hs' : s ≠ 1 := fun h ↦ by simp only [h, one_re, lt_self_iff_false] at hs - simp only [ne_eq, hs', not_false_eq_true, Function.update_noteq, B.χ.LFunction_eq_LSeries hs] - congr 1 - · simp_rw [← LSeries_zeta_eq_riemannZeta hs, ← natCoe_apply] - · exact LSeries_congr s B.χ.apply_eq_toArithmeticFunction_apply - -- summability side goals from `LSeries_convolution'` - · exact LSeriesSummable_zeta_iff.mpr hs - · exact (LSeriesSummable_congr _ fun h ↦ (B.χ.apply_eq_toArithmeticFunction_apply h).symm).mpr <| - ZMod.LSeriesSummable_of_one_lt_re B.χ hs - -/-- If `χ` is a bad character, then `F` is an entire function. -/ -private lemma F_differentiable (B : BadChar N) : Differentiable ℂ B.F := by - intro s - rcases ne_or_eq s 1 with hs | rfl - · exact B.F_differentiableAt_of_ne hs - -- now need to deal with `s = 1` - refine (analyticAt_of_differentiable_on_punctured_nhds_of_continuousAt ?_ ?_).differentiableAt - · filter_upwards [self_mem_nhdsWithin] with t ht - exact B.F_differentiableAt_of_ne ht - -- now reduced to showing *continuity* at s = 1 - let G := Function.update (fun s ↦ (s - 1) * riemannZeta s) 1 1 - let H := Function.update (fun s ↦ (B.χ.LFunction s - B.χ.LFunction 1) / (s - 1)) 1 - (deriv B.χ.LFunction 1) - have : B.F = G * H := by - ext1 t - rcases eq_or_ne t 1 with rfl | ht - · simp only [F, G, H, Pi.mul_apply, one_mul, Function.update_same] - · simp only [F, G, H, Function.update_noteq ht, mul_comm _ (riemannZeta _), B.hχ, sub_zero, - Pi.mul_apply, mul_assoc, mul_div_cancel₀ _ (sub_ne_zero.mpr ht)] - rw [this] - apply ContinuousAt.mul - · simpa only [G, continuousAt_update_same] using riemannZeta_residue_one - · exact (B.χ.differentiableAt_LFunction 1 (.inr B.χ_ne)).hasDerivAt.continuousAt_div - -/-- The trivial zero at `s = -2` of the zeta function gives that `F (-2) = 0`. -This is used later to obtain a contradction. -/ -private lemma F_neg_two (B : BadChar N) : B.F (-2 : ℝ) = 0 := by - have := riemannZeta_neg_two_mul_nat_add_one 0 - rw [Nat.cast_zero, zero_add, mul_one] at this - rw [F, ofReal_neg, ofReal_ofNat, Function.update_noteq (mod_cast (by omega : (-2 : ℤ) ≠ 1)), - this, zero_mul] - -end BadChar - -/-! -### The main result --/ - -/-- If `χ` is a nontrivial quadratic Dirichlet character, then `L(χ, 1) ≠ 0`. -/ -theorem LFunction_at_one_ne_zero_of_quadratic {N : ℕ} [NeZero N] {χ : DirichletCharacter ℂ N} - (hχ : χ ^ 2 = 1) (χ_ne : χ ≠ 1) : - χ.LFunction 1 ≠ 0 := by - intro hL - -- construct a "bad character" and put together a contradiction. - let B : BadChar N := {χ := χ, χ_sq := hχ, hχ := hL, χ_ne := χ_ne} - refine B.F_neg_two.not_gt ?_ - refine ArithmeticFunction.LSeries_positive_of_differentiable_of_eqOn (zetaMul_nonneg hχ) - (χ.isMultiplicative_zetaMul.map_one ▸ zero_lt_one) B.F_differentiable ?_ - (fun _ ↦ B.F_eq_LSeries) _ - exact LSeries.abscissaOfAbsConv_le_of_forall_lt_LSeriesSummable - fun _ a ↦ χ.LSeriesSummable_zetaMul a - -end DirichletCharacter From db944bfdc305d904efb85a54481226d2b7ee10d6 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 18 Nov 2024 11:19:55 +0000 Subject: [PATCH 059/829] chore: use `Pairwise (_ on _)` where possible (#9256) Co-authored-by: Jireh Loreaux Co-authored-by: Yury G. Kudryashov --- .../InnerProductSpace/JointEigenspace.lean | 4 ++-- Mathlib/Analysis/Normed/Algebra/Exponential.lean | 2 +- .../Normed/Algebra/MatrixExponential.lean | 2 +- Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean | 2 +- Mathlib/Data/Finset/NoncommProd.lean | 12 ++++++------ Mathlib/Data/Fintype/Basic.lean | 5 +++-- Mathlib/Data/List/Nodup.lean | 2 +- Mathlib/Data/List/Pairwise.lean | 2 +- Mathlib/Data/List/Permutation.lean | 2 +- Mathlib/Data/List/Sublists.lean | 2 +- Mathlib/Data/Multiset/Bind.lean | 6 ++++-- Mathlib/Data/Nat/Choose/Multinomial.lean | 4 ++-- Mathlib/Data/Set/Lattice.lean | 6 +++--- Mathlib/Data/Setoid/Partition.lean | 4 ++-- Mathlib/Data/ZMod/Quotient.lean | 2 +- Mathlib/GroupTheory/CoprodI.lean | 8 ++++---- Mathlib/Probability/StrongLaw.lean | 15 ++++++++------- Mathlib/RingTheory/Ideal/Quotient/Operations.lean | 6 +++--- Mathlib/SetTheory/Cardinal/Basic.lean | 4 ++-- 19 files changed, 47 insertions(+), 43 deletions(-) diff --git a/Mathlib/Analysis/InnerProductSpace/JointEigenspace.lean b/Mathlib/Analysis/InnerProductSpace/JointEigenspace.lean index 1502a88f847f7..95ac150524bfc 100644 --- a/Mathlib/Analysis/InnerProductSpace/JointEigenspace.lean +++ b/Mathlib/Analysis/InnerProductSpace/JointEigenspace.lean @@ -114,7 +114,7 @@ theorem directSum_isInternal_of_commute (hA : A.IsSymmetric) (hB : B.IsSymmetric /-- A commuting family of symmetric linear maps on a finite dimensional inner product space is simultaneously diagonalizable. -/ theorem iSup_iInf_eq_top_of_commute {ι : Type*} {T : ι → E →ₗ[𝕜] E} - (hT : ∀ i, (T i).IsSymmetric) (h : Pairwise fun i j ↦ Commute (T i) (T j)): + (hT : ∀ i, (T i).IsSymmetric) (h : Pairwise (Commute on T)): ⨆ χ : ι → 𝕜, ⨅ i, eigenspace (T i) (χ i) = ⊤ := calc _ = ⨆ χ : ι → 𝕜, ⨅ i, maxGenEigenspace (T i) (χ i) := @@ -129,7 +129,7 @@ theorem iSup_iInf_eq_top_of_commute {ι : Type*} {T : ι → E →ₗ[𝕜] E} /-- In finite dimensions, given a commuting family of symmetric linear operators, the inner product space on which they act decomposes as an internal direct sum of joint eigenspaces. -/ theorem LinearMap.IsSymmetric.directSum_isInternal_of_pairwise_commute [DecidableEq (n → 𝕜)] - (hT : ∀ i, (T i).IsSymmetric) (hC : Pairwise fun i j ↦ Commute (T i) (T j)) : + (hT : ∀ i, (T i).IsSymmetric) (hC : Pairwise (Commute on T)) : DirectSum.IsInternal (fun α : n → 𝕜 ↦ ⨅ j, eigenspace (T j) (α j)) := by rw [OrthogonalFamily.isInternal_iff] · rw [iSup_iInf_eq_top_of_commute hT hC, top_orthogonal_eq_bot] diff --git a/Mathlib/Analysis/Normed/Algebra/Exponential.lean b/Mathlib/Analysis/Normed/Algebra/Exponential.lean index 9f85562279b0c..2c3861688440d 100644 --- a/Mathlib/Analysis/Normed/Algebra/Exponential.lean +++ b/Mathlib/Analysis/Normed/Algebra/Exponential.lean @@ -477,7 +477,7 @@ end /-- In a Banach-algebra `𝔸` over `𝕂 = ℝ` or `𝕂 = ℂ`, if a family of elements `f i` mutually commute then `NormedSpace.exp 𝕂 (∑ i, f i) = ∏ i, NormedSpace.exp 𝕂 (f i)`. -/ theorem exp_sum_of_commute {ι} (s : Finset ι) (f : ι → 𝔸) - (h : (s : Set ι).Pairwise fun i j => Commute (f i) (f j)) : + (h : (s : Set ι).Pairwise (Commute on f)) : exp 𝕂 (∑ i ∈ s, f i) = s.noncommProd (fun i => exp 𝕂 (f i)) fun _ hi _ hj _ => (h.of_refl hi hj).exp 𝕂 := by classical diff --git a/Mathlib/Analysis/Normed/Algebra/MatrixExponential.lean b/Mathlib/Analysis/Normed/Algebra/MatrixExponential.lean index 497cbcfabba71..fe33c719a4832 100644 --- a/Mathlib/Analysis/Normed/Algebra/MatrixExponential.lean +++ b/Mathlib/Analysis/Normed/Algebra/MatrixExponential.lean @@ -125,7 +125,7 @@ nonrec theorem exp_add_of_commute (A B : Matrix m m 𝔸) (h : Commute A B) : exact exp_add_of_commute h nonrec theorem exp_sum_of_commute {ι} (s : Finset ι) (f : ι → Matrix m m 𝔸) - (h : (s : Set ι).Pairwise fun i j => Commute (f i) (f j)) : + (h : (s : Set ι).Pairwise (Commute on f)) : exp 𝕂 (∑ i ∈ s, f i) = s.noncommProd (fun i => exp 𝕂 (f i)) fun _ hi _ hj _ => (h.of_refl hi hj).exp 𝕂 := by letI : SeminormedRing (Matrix m m 𝔸) := Matrix.linftyOpSemiNormedRing diff --git a/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean b/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean index a3608131c3044..51f860e725e71 100644 --- a/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean +++ b/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean @@ -107,7 +107,7 @@ theorem nodup_antidiagonalTuple (k n : ℕ) : List.Nodup (antidiagonalTuple k n) dsimp at h₁₂ rw [Fin.cons_eq_cons, Nat.succ_inj'] at h₁₂ obtain ⟨h₁₂, rfl⟩ := h₁₂ - rw [h₁₂] at h + rw [Function.onFun, h₁₂] at h exact h (List.mem_map_of_mem _ hx₁) (List.mem_map_of_mem _ hx₂) theorem antidiagonalTuple_zero_right : ∀ k, antidiagonalTuple k 0 = [0] diff --git a/Mathlib/Data/Finset/NoncommProd.lean b/Mathlib/Data/Finset/NoncommProd.lean index 7140198dd8237..712c62a411ec7 100644 --- a/Mathlib/Data/Finset/NoncommProd.lean +++ b/Mathlib/Data/Finset/NoncommProd.lean @@ -232,7 +232,7 @@ variable [Monoid β] [Monoid γ] /-- Proof used in definition of `Finset.noncommProd` -/ @[to_additive] theorem noncommProd_lemma (s : Finset α) (f : α → β) - (comm : (s : Set α).Pairwise fun a b => Commute (f a) (f b)) : + (comm : (s : Set α).Pairwise (Commute on f)) : Set.Pairwise { x | x ∈ Multiset.map f s.val } Commute := by simp_rw [Multiset.mem_map] rintro _ ⟨a, ha, rfl⟩ _ ⟨b, hb, rfl⟩ _ @@ -244,7 +244,7 @@ given a proof that `*` commutes on all elements `f x` for `x ∈ s`. -/ "Sum of a `s : Finset α` mapped with `f : α → β` with `[AddMonoid β]`, given a proof that `+` commutes on all elements `f x` for `x ∈ s`."] def noncommProd (s : Finset α) (f : α → β) - (comm : (s : Set α).Pairwise fun a b => Commute (f a) (f b)) : β := + (comm : (s : Set α).Pairwise (Commute on f)) : β := (s.1.map f).noncommProd <| noncommProd_lemma s f comm @[to_additive] @@ -260,7 +260,7 @@ theorem noncommProd_congr {s₁ s₂ : Finset α} {f g : α → β} (h₁ : s₁ (h₂ : ∀ x ∈ s₂, f x = g x) (comm) : noncommProd s₁ f comm = noncommProd s₂ g fun x hx y hy h => by - dsimp only + dsimp only [Function.onFun] rw [← h₂ _ hx, ← h₂ _ hy] subst h₁ exact comm hx hy h := by @@ -363,7 +363,7 @@ theorem noncommProd_eq_prod {β : Type*} [CommMonoid β] (s : Finset α) (f : α /-- The non-commutative version of `Finset.prod_union` -/ @[to_additive "The non-commutative version of `Finset.sum_union`"] theorem noncommProd_union_of_disjoint [DecidableEq α] {s t : Finset α} (h : Disjoint s t) - (f : α → β) (comm : { x | x ∈ s ∪ t }.Pairwise fun a b => Commute (f a) (f b)) : + (f : α → β) (comm : { x | x ∈ s ∪ t }.Pairwise (Commute on f)) : noncommProd (s ∪ t) f comm = noncommProd s f (comm.mono <| coe_subset.2 subset_union_left) * noncommProd t f (comm.mono <| coe_subset.2 subset_union_right) := by @@ -379,8 +379,8 @@ theorem noncommProd_union_of_disjoint [DecidableEq α] {s t : Finset α} (h : Di @[to_additive] theorem noncommProd_mul_distrib_aux {s : Finset α} {f : α → β} {g : α → β} - (comm_ff : (s : Set α).Pairwise fun x y => Commute (f x) (f y)) - (comm_gg : (s : Set α).Pairwise fun x y => Commute (g x) (g y)) + (comm_ff : (s : Set α).Pairwise (Commute on f)) + (comm_gg : (s : Set α).Pairwise (Commute on g)) (comm_gf : (s : Set α).Pairwise fun x y => Commute (g x) (f y)) : (s : Set α).Pairwise fun x y => Commute ((f * g) x) ((f * g) y) := by intro x hx y hy h diff --git a/Mathlib/Data/Fintype/Basic.lean b/Mathlib/Data/Fintype/Basic.lean index 0f5c6fcb5fe92..90e35e73ab619 100644 --- a/Mathlib/Data/Fintype/Basic.lean +++ b/Mathlib/Data/Fintype/Basic.lean @@ -1174,13 +1174,14 @@ function `f : ℕ → α` such that `r (f m) (f n)` holds whenever `m ≠ n`. We also ensure that all constructed points satisfy a given predicate `P`. -/ theorem exists_seq_of_forall_finset_exists' {α : Type*} (P : α → Prop) (r : α → α → Prop) [IsSymm α r] (h : ∀ s : Finset α, (∀ x ∈ s, P x) → ∃ y, P y ∧ ∀ x ∈ s, r x y) : - ∃ f : ℕ → α, (∀ n, P (f n)) ∧ Pairwise fun m n => r (f m) (f n) := by + ∃ f : ℕ → α, (∀ n, P (f n)) ∧ Pairwise (r on f) := by rcases exists_seq_of_forall_finset_exists P r h with ⟨f, hf, hf'⟩ refine ⟨f, hf, fun m n hmn => ?_⟩ rcases lt_trichotomy m n with (h | rfl | h) · exact hf' m n h · exact (hmn rfl).elim - · apply symm + · unfold Function.onFun + apply symm exact hf' n m h open Batteries.ExtendedBinder Lean Meta diff --git a/Mathlib/Data/List/Nodup.lean b/Mathlib/Data/List/Nodup.lean index 26f193504d547..f6b34a9b86b77 100644 --- a/Mathlib/Data/List/Nodup.lean +++ b/Mathlib/Data/List/Nodup.lean @@ -259,7 +259,7 @@ theorem nodup_flatten {L : List (List α)} : theorem nodup_flatMap {l₁ : List α} {f : α → List β} : Nodup (l₁.flatMap f) ↔ - (∀ x ∈ l₁, Nodup (f x)) ∧ Pairwise (fun a b : α => Disjoint (f a) (f b)) l₁ := by + (∀ x ∈ l₁, Nodup (f x)) ∧ Pairwise (Disjoint on f) l₁ := by simp only [List.flatMap, nodup_flatten, pairwise_map, and_comm, and_left_comm, mem_map, exists_imp, and_imp] rw [show (∀ (l : List β) (x : α), f x = l → x ∈ l₁ → Nodup l) ↔ ∀ x : α, x ∈ l₁ → Nodup (f x) diff --git a/Mathlib/Data/List/Pairwise.lean b/Mathlib/Data/List/Pairwise.lean index cb395fad73168..4682916a610a0 100644 --- a/Mathlib/Data/List/Pairwise.lean +++ b/Mathlib/Data/List/Pairwise.lean @@ -51,7 +51,7 @@ theorem Pairwise.set_pairwise (hl : Pairwise R l) (hr : Symmetric R) : { x | x -- Porting note: Duplicate of `pairwise_map` but with `f` explicit. @[deprecated (since := "2024-02-25")] theorem pairwise_map' (f : β → α) : - ∀ {l : List β}, Pairwise R (map f l) ↔ Pairwise (fun a b : β => R (f a) (f b)) l + ∀ {l : List β}, Pairwise R (map f l) ↔ Pairwise (R on f) l | [] => by simp only [map, Pairwise.nil] | b :: l => by simp only [map, pairwise_cons, mem_map, forall_exists_index, and_imp, diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index 2ea051630ee09..d90a8bbe313e9 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -484,7 +484,7 @@ theorem nodup_permutations (s : List α) (hs : Nodup s) : Nodup s.permutations : rw [nodup_permutations'Aux_iff, hy.mem_iff] exact fun H => h x H rfl · refine IH.pairwise_of_forall_ne fun as ha bs hb H => ?_ - rw [disjoint_iff_ne] + rw [Function.onFun, disjoint_iff_ne] rintro a ha' b hb' rfl obtain ⟨⟨n, hn⟩, hn'⟩ := get_of_mem ha' obtain ⟨⟨m, hm⟩, hm'⟩ := get_of_mem hb' diff --git a/Mathlib/Data/List/Sublists.lean b/Mathlib/Data/List/Sublists.lean index 9659c779dff54..cc89adca8c9e8 100644 --- a/Mathlib/Data/List/Sublists.lean +++ b/Mathlib/Data/List/Sublists.lean @@ -317,7 +317,7 @@ theorem Pairwise.sublists' {R} : exact Lex.rel (H₁ _ <| sl₁.subset <| mem_cons_self _ _) theorem pairwise_sublists {R} {l : List α} (H : Pairwise R l) : - Pairwise (fun l₁ l₂ => Lex R (reverse l₁) (reverse l₂)) (sublists l) := by + Pairwise (Lex R on reverse) (sublists l) := by have := (pairwise_reverse.2 H).sublists' rwa [sublists'_reverse, pairwise_map] at this diff --git a/Mathlib/Data/Multiset/Bind.lean b/Mathlib/Data/Multiset/Bind.lean index 4419499efba7e..b34913bef0ace 100644 --- a/Mathlib/Data/Multiset/Bind.lean +++ b/Mathlib/Data/Multiset/Bind.lean @@ -202,12 +202,14 @@ theorem attach_bind_coe (s : Multiset α) (f : α → Multiset β) : variable {f s t} @[simp] lemma nodup_bind : - Nodup (bind s f) ↔ (∀ a ∈ s, Nodup (f a)) ∧ s.Pairwise fun a b => Disjoint (f a) (f b) := by + Nodup (bind s f) ↔ (∀ a ∈ s, Nodup (f a)) ∧ s.Pairwise (Disjoint on f) := by have : ∀ a, ∃ l : List β, f a = l := fun a => Quot.induction_on (f a) fun l => ⟨l, rfl⟩ choose f' h' using this have : f = fun a ↦ ofList (f' a) := funext h' have hd : Symmetric fun a b ↦ List.Disjoint (f' a) (f' b) := fun a b h ↦ h.symm - exact Quot.induction_on s <| by simp [this, List.nodup_flatMap, pairwise_coe_iff_pairwise hd] + exact Quot.induction_on s <| by + unfold Function.onFun + simp [this, List.nodup_flatMap, pairwise_coe_iff_pairwise hd] @[simp] lemma dedup_bind_dedup [DecidableEq α] [DecidableEq β] (s : Multiset α) (f : α → Multiset β) : diff --git a/Mathlib/Data/Nat/Choose/Multinomial.lean b/Mathlib/Data/Nat/Choose/Multinomial.lean index 19e15e9bc63ee..51ba07f30c0c7 100644 --- a/Mathlib/Data/Nat/Choose/Multinomial.lean +++ b/Mathlib/Data/Nat/Choose/Multinomial.lean @@ -210,7 +210,7 @@ variable [Semiring R] -- TODO: Can we prove one of the following two from the other one? /-- The **multinomial theorem**. -/ lemma sum_pow_eq_sum_piAntidiag_of_commute (s : Finset α) (f : α → R) - (hc : (s : Set α).Pairwise fun i j ↦ Commute (f i) (f j)) (n : ℕ) : + (hc : (s : Set α).Pairwise (Commute on f)) (n : ℕ) : (∑ i in s, f i) ^ n = ∑ k in piAntidiag s n, multinomial s k * s.noncommProd (fun i ↦ f i ^ k i) (hc.mono' fun _ _ h ↦ h.pow_pow ..) := by classical @@ -250,7 +250,7 @@ lemma sum_pow_eq_sum_piAntidiag_of_commute (s : Finset α) (f : α → R) /-- The **multinomial theorem**. -/ theorem sum_pow_of_commute (x : α → R) (s : Finset α) - (hc : (s : Set α).Pairwise fun i j => Commute (x i) (x j)) : + (hc : (s : Set α).Pairwise (Commute on x)) : ∀ n, s.sum x ^ n = ∑ k : s.sym n, diff --git a/Mathlib/Data/Set/Lattice.lean b/Mathlib/Data/Set/Lattice.lean index 089ca2b1ee2cd..41676f753e553 100644 --- a/Mathlib/Data/Set/Lattice.lean +++ b/Mathlib/Data/Set/Lattice.lean @@ -1893,7 +1893,7 @@ theorem sigmaToiUnion_surjective : Surjective (sigmaToiUnion t) let ⟨a, hb⟩ := this ⟨⟨a, b, hb⟩, rfl⟩ -theorem sigmaToiUnion_injective (h : Pairwise fun i j => Disjoint (t i) (t j)) : +theorem sigmaToiUnion_injective (h : Pairwise (Disjoint on t)) : Injective (sigmaToiUnion t) | ⟨a₁, b₁, h₁⟩, ⟨a₂, b₂, h₂⟩, eq => have b_eq : b₁ = b₂ := congr_arg Subtype.val eq @@ -1903,7 +1903,7 @@ theorem sigmaToiUnion_injective (h : Pairwise fun i j => Disjoint (t i) (t j)) : (h ne).le_bot this Sigma.eq a_eq <| Subtype.eq <| by subst b_eq; subst a_eq; rfl -theorem sigmaToiUnion_bijective (h : Pairwise fun i j => Disjoint (t i) (t j)) : +theorem sigmaToiUnion_bijective (h : Pairwise (Disjoint on t)) : Bijective (sigmaToiUnion t) := ⟨sigmaToiUnion_injective t h, sigmaToiUnion_surjective t⟩ @@ -1918,7 +1918,7 @@ noncomputable def sigmaEquiv (s : α → Set β) (hs : ∀ b, ∃! i, b ∈ s i) /-- Equivalence between a disjoint union and a dependent sum. -/ noncomputable def unionEqSigmaOfDisjoint {t : α → Set β} - (h : Pairwise fun i j => Disjoint (t i) (t j)) : + (h : Pairwise (Disjoint on t)) : (⋃ i, t i) ≃ Σi, t i := (Equiv.ofBijective _ <| sigmaToiUnion_bijective t h).symm diff --git a/Mathlib/Data/Setoid/Partition.lean b/Mathlib/Data/Setoid/Partition.lean index c362e4ce974c2..92f2964a06ba7 100644 --- a/Mathlib/Data/Setoid/Partition.lean +++ b/Mathlib/Data/Setoid/Partition.lean @@ -312,7 +312,7 @@ structure IndexedPartition {ι α : Type*} (s : ι → Set α) where /-- The non-constructive constructor for `IndexedPartition`. -/ noncomputable def IndexedPartition.mk' {ι α : Type*} (s : ι → Set α) - (dis : Pairwise fun i j => Disjoint (s i) (s j)) (nonempty : ∀ i, (s i).Nonempty) + (dis : Pairwise (Disjoint on s)) (nonempty : ∀ i, (s i).Nonempty) (ex : ∀ x, ∃ i, x ∈ s i) : IndexedPartition s where eq_of_mem {_x _i _j} hxi hxj := by_contradiction fun h => (dis h).le_bot ⟨hxi, hxj⟩ some i := (nonempty i).some @@ -349,7 +349,7 @@ theorem iUnion : ⋃ i, s i = univ := by simp [hs.exists_mem x] include hs in -theorem disjoint : Pairwise fun i j => Disjoint (s i) (s j) := fun {_i _j} h => +theorem disjoint : Pairwise (Disjoint on s) := fun {_i _j} h => disjoint_left.mpr fun {_x} hxi hxj => h (hs.eq_of_mem hxi hxj) theorem mem_iff_index_eq {x i} : x ∈ s i ↔ hs.index x = i := diff --git a/Mathlib/Data/ZMod/Quotient.lean b/Mathlib/Data/ZMod/Quotient.lean index 031fa4b2b17ad..1e52c60f93728 100644 --- a/Mathlib/Data/ZMod/Quotient.lean +++ b/Mathlib/Data/ZMod/Quotient.lean @@ -69,7 +69,7 @@ open Ideal /-- The **Chinese remainder theorem**, elementary version for `ZMod`. See also `Mathlib.Data.ZMod.Basic` for versions involving only two numbers. -/ def ZMod.prodEquivPi {ι : Type*} [Fintype ι] (a : ι → ℕ) - (coprime : Pairwise fun i j => Nat.Coprime (a i) (a j)) : ZMod (∏ i, a i) ≃+* Π i, ZMod (a i) := + (coprime : Pairwise (Nat.Coprime on a)) : ZMod (∏ i, a i) ≃+* Π i, ZMod (a i) := have : Pairwise fun i j => IsCoprime (span {(a i : ℤ)}) (span {(a j : ℤ)}) := fun _i _j h ↦ (isCoprime_span_singleton_iff _ _).mpr ((coprime h).cast (R := ℤ)) Int.quotientSpanNatEquivZMod _ |>.symm.trans <| diff --git a/Mathlib/GroupTheory/CoprodI.lean b/Mathlib/GroupTheory/CoprodI.lean index 1a41a92b0ff56..5e020fad4e35f 100644 --- a/Mathlib/GroupTheory/CoprodI.lean +++ b/Mathlib/GroupTheory/CoprodI.lean @@ -823,7 +823,7 @@ variable (hcard : 3 ≤ #ι ∨ ∃ i, 3 ≤ #(H i)) variable {α : Type*} [MulAction G α] variable (X : ι → Set α) variable (hXnonempty : ∀ i, (X i).Nonempty) -variable (hXdisj : Pairwise fun i j => Disjoint (X i) (X j)) +variable (hXdisj : Pairwise (Disjoint on X)) variable (hpp : Pairwise fun i j => ∀ h : H i, h ≠ 1 → f i h • X j ⊆ X i) include hpp @@ -973,8 +973,8 @@ variable {G : Type u_1} [Group G] (a : ι → G) variable {α : Type*} [MulAction G α] variable (X Y : ι → Set α) variable (hXnonempty : ∀ i, (X i).Nonempty) -variable (hXdisj : Pairwise fun i j => Disjoint (X i) (X j)) -variable (hYdisj : Pairwise fun i j => Disjoint (Y i) (Y j)) +variable (hXdisj : Pairwise (Disjoint on X)) +variable (hYdisj : Pairwise (Disjoint on Y)) variable (hXYdisj : ∀ i j, Disjoint (X i) (Y j)) variable (hX : ∀ i, a i • (Y i)ᶜ ⊆ X i) variable (hY : ∀ i, a⁻¹ i • (X i)ᶜ ⊆ Y i) @@ -1009,7 +1009,7 @@ theorem _root_.FreeGroup.injective_lift_of_ping_pong : Function.Injective (FreeG apply lift_injective_of_ping_pong f _ X' · show ∀ i, (X' i).Nonempty exact fun i => Set.Nonempty.inl (hXnonempty i) - · show Pairwise fun i j => Disjoint (X' i) (X' j) + · show Pairwise (Disjoint on X') intro i j hij simp only [X'] apply Disjoint.union_left <;> apply Disjoint.union_right diff --git a/Mathlib/Probability/StrongLaw.lean b/Mathlib/Probability/StrongLaw.lean index 4158b48564a82..4740c2314e1fb 100644 --- a/Mathlib/Probability/StrongLaw.lean +++ b/Mathlib/Probability/StrongLaw.lean @@ -369,7 +369,7 @@ pairwise independence) for nonnegative random variables, following Etemadi's pro section StrongLawNonneg variable (X : ℕ → Ω → ℝ) (hint : Integrable (X 0)) - (hindep : Pairwise fun i j => IndepFun (X i) (X j)) (hident : ∀ i, IdentDistrib (X i) (X 0)) + (hindep : Pairwise (IndepFun on X)) (hident : ∀ i, IdentDistrib (X i) (X 0)) (hnonneg : ∀ i ω, 0 ≤ X i ω) include hint hindep hident hnonneg in @@ -604,7 +604,7 @@ requires pairwise independence. Superseded by `strong_law_ae`, which works for r taking values in any Banach space. -/ theorem strong_law_ae_real {Ω : Type*} {m : MeasurableSpace Ω} {μ : Measure Ω} (X : ℕ → Ω → ℝ) (hint : Integrable (X 0) μ) - (hindep : Pairwise fun i j => IndepFun (X i) (X j) μ) + (hindep : Pairwise ((IndepFun · · μ) on X)) (hident : ∀ i, IdentDistrib (X i) (X 0) μ μ) : ∀ᵐ ω ∂μ, Tendsto (fun n : ℕ => (∑ i ∈ range n, X i ω) / n) atTop (𝓝 (μ [X 0])) := by let mΩ : MeasureSpace Ω := ⟨μ⟩ @@ -651,7 +651,7 @@ open Set TopologicalSpace the composition of the random variables with a simple function satisfies the strong law of large numbers. -/ lemma strong_law_ae_simpleFunc_comp (X : ℕ → Ω → E) (h' : Measurable (X 0)) - (hindep : Pairwise (fun i j ↦ IndepFun (X i) (X j) μ)) + (hindep : Pairwise ((IndepFun · · μ) on X)) (hident : ∀ i, IdentDistrib (X i) (X 0) μ μ) (φ : SimpleFunc E E) : ∀ᵐ ω ∂μ, Tendsto (fun n : ℕ ↦ (n : ℝ) ⁻¹ • (∑ i ∈ range n, φ (X i ω))) atTop (𝓝 (μ [φ ∘ (X 0)])) := by @@ -702,7 +702,7 @@ assuming measurability in addition to integrability. This is weakened to ae meas the full version `ProbabilityTheory.strong_law_ae`. -/ lemma strong_law_ae_of_measurable (X : ℕ → Ω → E) (hint : Integrable (X 0) μ) (h' : StronglyMeasurable (X 0)) - (hindep : Pairwise (fun i j ↦ IndepFun (X i) (X j) μ)) + (hindep : Pairwise ((IndepFun · · μ) on X)) (hident : ∀ i, IdentDistrib (X i) (X 0) μ μ) : ∀ᵐ ω ∂μ, Tendsto (fun n : ℕ ↦ (n : ℝ) ⁻¹ • (∑ i ∈ range n, X i ω)) atTop (𝓝 (μ [X 0])) := by /- Choose a simple function `φ` such that `φ (X 0)` approximates well enough `X 0` -- this is @@ -733,7 +733,8 @@ lemma strong_law_ae_of_measurable have I : ∀ k i, (fun ω ↦ ‖(X i - Y k i) ω‖) = (G k) ∘ (X i) := fun k i ↦ rfl apply strong_law_ae_real (fun i ω ↦ ‖(X i - Y k i) ω‖) · exact (hint.sub ((φ k).comp (X 0) h'.measurable).integrable_of_isFiniteMeasure).norm - · simp_rw [I] + · unfold Function.onFun + simp_rw [I] intro i j hij exact (hindep hij).comp (G_meas k) (G_meas k) · intro i @@ -792,7 +793,7 @@ identically distributed integrable random variables taking values in a Banach sp then `n⁻¹ • ∑ i ∈ range n, X i` converges almost surely to `𝔼[X 0]`. We give here the strong version, due to Etemadi, that only requires pairwise independence. -/ theorem strong_law_ae (X : ℕ → Ω → E) (hint : Integrable (X 0) μ) - (hindep : Pairwise (fun i j ↦ IndepFun (X i) (X j) μ)) + (hindep : Pairwise ((IndepFun · · μ) on X)) (hident : ∀ i, IdentDistrib (X i) (X 0) μ μ) : ∀ᵐ ω ∂μ, Tendsto (fun n : ℕ ↦ (n : ℝ) ⁻¹ • (∑ i ∈ range n, X i ω)) atTop (𝓝 (μ [X 0])) := by -- First exclude the trivial case where the space is not a probability space @@ -836,7 +837,7 @@ variable {Ω : Type*} {mΩ : MeasurableSpace Ω} {μ : Measure Ω} identically distributed random variables in Lᵖ, then `n⁻¹ • ∑ i ∈ range n, X i` converges in `Lᵖ` to `𝔼[X 0]`. -/ theorem strong_law_Lp {p : ℝ≥0∞} (hp : 1 ≤ p) (hp' : p ≠ ∞) (X : ℕ → Ω → E) - (hℒp : Memℒp (X 0) p μ) (hindep : Pairwise fun i j => IndepFun (X i) (X j) μ) + (hℒp : Memℒp (X 0) p μ) (hindep : Pairwise ((IndepFun · · μ) on X)) (hident : ∀ i, IdentDistrib (X i) (X 0) μ μ) : Tendsto (fun (n : ℕ) => eLpNorm (fun ω => (n : ℝ) ⁻¹ • (∑ i ∈ range n, X i ω) - μ [X 0]) p μ) atTop (𝓝 0) := by diff --git a/Mathlib/RingTheory/Ideal/Quotient/Operations.lean b/Mathlib/RingTheory/Ideal/Quotient/Operations.lean index 76b15f82bbce4..8d8291b64c232 100644 --- a/Mathlib/RingTheory/Ideal/Quotient/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Quotient/Operations.lean @@ -198,7 +198,7 @@ lemma quotientInfToPiQuotient_inj (I : ι → Ideal R) : Injective (quotientInfT rw [quotientInfToPiQuotient, injective_lift_iff, ker_Pi_Quotient_mk] lemma quotientInfToPiQuotient_surj [Finite ι] {I : ι → Ideal R} - (hI : Pairwise fun i j => IsCoprime (I i) (I j)) : Surjective (quotientInfToPiQuotient I) := by + (hI : Pairwise (IsCoprime on I)) : Surjective (quotientInfToPiQuotient I) := by classical cases nonempty_fintype ι intro g @@ -224,7 +224,7 @@ lemma quotientInfToPiQuotient_surj [Finite ι] {I : ι → Ideal R} /-- **Chinese Remainder Theorem**. Eisenbud Ex.2.6. Similar to Atiyah-Macdonald 1.10 and Stacks 00DT -/ noncomputable def quotientInfRingEquivPiQuotient [Finite ι] (f : ι → Ideal R) - (hf : Pairwise fun i j => IsCoprime (f i) (f j)) : (R ⧸ ⨅ i, f i) ≃+* ∀ i, R ⧸ f i := + (hf : Pairwise (IsCoprime on f)) : (R ⧸ ⨅ i, f i) ≃+* ∀ i, R ⧸ f i := { Equiv.ofBijective _ ⟨quotientInfToPiQuotient_inj f, quotientInfToPiQuotient_surj hf⟩, quotientInfToPiQuotient f with } @@ -250,7 +250,7 @@ lemma exists_forall_sub_mem_ideal {R : Type*} [CommRing R] {ι : Type*} [Finite noncomputable def quotientInfEquivQuotientProd (I J : Ideal R) (coprime : IsCoprime I J) : R ⧸ I ⊓ J ≃+* (R ⧸ I) × R ⧸ J := let f : Fin 2 → Ideal R := ![I, J] - have hf : Pairwise fun i j => IsCoprime (f i) (f j) := by + have hf : Pairwise (IsCoprime on f) := by intro i j h fin_cases i <;> fin_cases j <;> try contradiction · assumption diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index c3612fa0c43c8..40f07f1cc48f2 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -1880,13 +1880,13 @@ theorem mk_iUnion_le_sum_mk_lift {α : Type u} {ι : Type v} {f : ι → Set α} _ = sum fun i => #(f i) := mk_sigma _ theorem mk_iUnion_eq_sum_mk {α ι : Type u} {f : ι → Set α} - (h : Pairwise fun i j => Disjoint (f i) (f j)) : #(⋃ i, f i) = sum fun i => #(f i) := + (h : Pairwise (Disjoint on f)) : #(⋃ i, f i) = sum fun i => #(f i) := calc #(⋃ i, f i) = #(Σi, f i) := mk_congr (Set.unionEqSigmaOfDisjoint h) _ = sum fun i => #(f i) := mk_sigma _ theorem mk_iUnion_eq_sum_mk_lift {α : Type u} {ι : Type v} {f : ι → Set α} - (h : Pairwise fun i j => Disjoint (f i) (f j)) : + (h : Pairwise (Disjoint on f)) : lift.{v} #(⋃ i, f i) = sum fun i => #(f i) := calc lift.{v} #(⋃ i, f i) = #(Σi, f i) := From f2bb1c756349a7cfdf45fd4c1dc2399511981946 Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 18 Nov 2024 11:19:57 +0000 Subject: [PATCH 060/829] feat: the upstreamDecl linter (#18702) This linter is intended to work much like `minImports` in helping with splitting files. It places a warning on each set of declarations that could be moved to an earlier file. It takes into account dependencies on declarations in the current file, so that a whole block of declarations can be moved even if there are dependencies between the declarations in the block. Since definitions are often deliberately placed in their own file, this linter does not place a warning on any declaration depending on a definition in the current file. Co-authored-by: Anne C.A. Baanen Co-authored-by: Anne Baanen --- Mathlib.lean | 1 + Mathlib/Tactic.lean | 1 + Mathlib/Tactic/Linter.lean | 1 + Mathlib/Tactic/Linter/UpstreamableDecl.lean | 90 +++++++++++++++++++++ Mathlib/Tactic/MinImports.lean | 27 ++++++- MathlibTest/MinImports.lean | 35 ++++++++ 6 files changed, 151 insertions(+), 4 deletions(-) create mode 100644 Mathlib/Tactic/Linter/UpstreamableDecl.lean diff --git a/Mathlib.lean b/Mathlib.lean index e55b078c41c8f..d1d7a4d8d578f 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4697,6 +4697,7 @@ import Mathlib.Tactic.Linter.RefineLinter import Mathlib.Tactic.Linter.Style import Mathlib.Tactic.Linter.TextBased import Mathlib.Tactic.Linter.UnusedTactic +import Mathlib.Tactic.Linter.UpstreamableDecl import Mathlib.Tactic.Measurability import Mathlib.Tactic.Measurability.Init import Mathlib.Tactic.MinImports diff --git a/Mathlib/Tactic.lean b/Mathlib/Tactic.lean index fd01ee715c4c0..cbf02e204731c 100644 --- a/Mathlib/Tactic.lean +++ b/Mathlib/Tactic.lean @@ -152,6 +152,7 @@ import Mathlib.Tactic.Linter.RefineLinter import Mathlib.Tactic.Linter.Style import Mathlib.Tactic.Linter.TextBased import Mathlib.Tactic.Linter.UnusedTactic +import Mathlib.Tactic.Linter.UpstreamableDecl import Mathlib.Tactic.Measurability import Mathlib.Tactic.Measurability.Init import Mathlib.Tactic.MinImports diff --git a/Mathlib/Tactic/Linter.lean b/Mathlib/Tactic/Linter.lean index 2f67da1bea192..605f17197fb4a 100644 --- a/Mathlib/Tactic/Linter.lean +++ b/Mathlib/Tactic/Linter.lean @@ -12,3 +12,4 @@ import Mathlib.Tactic.Linter.FlexibleLinter import Mathlib.Tactic.Linter.HaveLetLinter import Mathlib.Tactic.Linter.MinImports import Mathlib.Tactic.Linter.PPRoundtrip +import Mathlib.Tactic.Linter.UpstreamableDecl diff --git a/Mathlib/Tactic/Linter/UpstreamableDecl.lean b/Mathlib/Tactic/Linter/UpstreamableDecl.lean new file mode 100644 index 0000000000000..86f55dc84e8bd --- /dev/null +++ b/Mathlib/Tactic/Linter/UpstreamableDecl.lean @@ -0,0 +1,90 @@ +/- +Copyright (c) 2024 Damiano Testa. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Damiano Testa, Anne Baanen +-/ +import ImportGraph.Imports +import Mathlib.Lean.Expr.Basic +import Mathlib.Tactic.MinImports + +/-! # The `upstreamableDecl` linter + +The `upstreamableDecl` linter detects declarations that could be moved to a file higher up in the +import hierarchy. This is intended to assist with splitting files. +-/ + +open Lean Elab Command + +/-- Does this declaration come from the current file? -/ +def Lean.Name.isLocal (env : Environment) (decl : Name) : Bool := + (env.getModuleIdxFor? decl).isNone + +open Mathlib.Command.MinImports + +/-- Does the declaration with this name depend on `def`s in the current file? -/ +def Lean.Environment.localDefDependencies (env : Environment) (stx id : Syntax) : + CommandElabM Bool := do + let declName : NameSet ← try + NameSet.ofList <$> resolveGlobalConst id + catch _ => + pure ∅ + + let immediateDeps ← getAllDependencies stx id + + -- Drop all the unresolvable constants, otherwise `transitivelyUsedConstants` fails. + let immediateDeps : NameSet := immediateDeps.fold (init := ∅) fun s n => + if (env.find? n).isSome then s.insert n else s + + let deps ← liftCoreM <| immediateDeps.transitivelyUsedConstants + let constInfos := deps.toList.filterMap env.find? + let defs := constInfos.filter (·.isDef) + return defs.any fun constInfo => !(declName.contains constInfo.name) && constInfo.name.isLocal env + +namespace Mathlib.Linter + +/-- +The `upstreamableDecl` linter detects declarations that could be moved to a file higher up in the +import hierarchy. If this is the case, it emits a warning. + +This is intended to assist with splitting files. + +The linter does not place a warning on any declaration depending on a definition in the current file +(while it does place a warning on the definition itself), since we often create a new file for a +definition on purpose. +-/ +register_option linter.upstreamableDecl : Bool := { + defValue := false + descr := "enable the upstreamableDecl linter" +} + +namespace DoubleImports + +@[inherit_doc Mathlib.Linter.linter.upstreamableDecl] +def upstreamableDeclLinter : Linter where run := withSetOptionIn fun stx ↦ do + unless Linter.getLinterValue linter.upstreamableDecl (← getOptions) do + return + if (← get).messages.hasErrors then + return + if stx == (← `(command| set_option $(mkIdent `linter.upstreamableDecl) true)) then return + let env ← getEnv + let id ← getId stx + if id != .missing then + let minImports := getIrredundantImports env (← getAllImports stx id) + match minImports with + | ⟨(RBNode.node _ .leaf upstream _ .leaf), _⟩ => do + if !(← env.localDefDependencies stx id) then + let p : GoToModuleLinkProps := { modName := upstream } + let widget : MessageData := .ofWidget + (← liftCoreM <| Widget.WidgetInstance.ofHash + GoToModuleLink.javascriptHash <| + Server.RpcEncodable.rpcEncode p) + (toString upstream) + Linter.logLint linter.upstreamableDecl id + m!"Consider moving this declaration to the module {widget}." + | _ => pure () + +initialize addLinter upstreamableDeclLinter + +end DoubleImports + +end Mathlib.Linter diff --git a/Mathlib/Tactic/MinImports.lean b/Mathlib/Tactic/MinImports.lean index f01c70978ae97..05f758440185c 100644 --- a/Mathlib/Tactic/MinImports.lean +++ b/Mathlib/Tactic/MinImports.lean @@ -131,8 +131,8 @@ def previousInstName : Name → Name .str init newTail | nm => nm -/--`getAllImports cmd id` takes a `Syntax` input `cmd` and returns the `NameSet` of all the -module names that are implied by +/--`getAllDependencies cmd id` takes a `Syntax` input `cmd` and returns the `NameSet` of all the +declaration names that are implied by * the `SyntaxNodeKinds`, * the attributes of `cmd` (if there are any), * the identifiers contained in `cmd`, @@ -140,8 +140,11 @@ module names that are implied by The argument `id` is expected to be an identifier. It is used either for the internally generated name of a "nameless" `instance` or when parsing an identifier representing the name of a declaration. + +Note that the return value does not contain dependencies of the dependencies; +you can use `Lean.NameSet.transitivelyUsedConstants` to get those. -/ -def getAllImports (cmd id : Syntax) (dbg? : Bool := false) : +def getAllDependencies (cmd id : Syntax) : CommandElabM NameSet := do let env ← getEnv let id1 ← getId cmd @@ -155,10 +158,26 @@ def getAllImports (cmd id : Syntax) (dbg? : Bool := false) : -- failing everything, use the current namespace followed by the visible name return ns ++ id1.getId) -- We collect the implied declaration names, the `SyntaxNodeKinds` and the attributes. - let ts := getVisited env nm + return getVisited env nm |>.append (getVisited env id.getId) |>.append (getSyntaxNodeKinds cmd) |>.append (getAttrs env cmd) + +/--`getAllImports cmd id` takes a `Syntax` input `cmd` and returns the `NameSet` of all the +module names that are implied by +* the `SyntaxNodeKinds`, +* the attributes of `cmd` (if there are any), +* the identifiers contained in `cmd`, +* if `cmd` adds a declaration `d` to the environment, then also all the module names implied by `d`. +The argument `id` is expected to be an identifier. +It is used either for the internally generated name of a "nameless" `instance` or when parsing +an identifier representing the name of a declaration. +-/ +def getAllImports (cmd id : Syntax) (dbg? : Bool := false) : + CommandElabM NameSet := do + let env ← getEnv + -- We collect the implied declaration names, the `SyntaxNodeKinds` and the attributes. + let ts ← getAllDependencies cmd id if dbg? then dbg_trace "{ts.toArray.qsort Name.lt}" let mut hm : Std.HashMap Nat Name := {} for imp in env.header.moduleNames do diff --git a/MathlibTest/MinImports.lean b/MathlibTest/MinImports.lean index d97b8e1cff7a6..43ee73fa011e9 100644 --- a/MathlibTest/MinImports.lean +++ b/MathlibTest/MinImports.lean @@ -74,6 +74,8 @@ import Mathlib.Data.Nat.Notation #min_imports in lemma hi (n : ℕ) : n = n := by extract_goal; rfl +section Linter.MinImports + set_option linter.minImports.increases false set_option linter.minImports true /-- @@ -131,3 +133,36 @@ note: this linter can be disabled with `set_option linter.minImports false` run_cmd let _ ← `(declModifiers|@[fun_prop]) let _ ← `(tactic|apply @Mathlib.Meta.NormNum.evalNatDvd <;> extract_goal) + +end Linter.MinImports + +section Linter.UpstreamableDecl + +set_option linter.upstreamableDecl true + +/-- +warning: Consider moving this declaration to the module Mathlib.Data.Nat.Notation. +note: this linter can be disabled with `set_option linter.upstreamableDecl false` +-/ +#guard_msgs in +theorem propose_to_move_this_theorem : (0 : ℕ) = 0 := rfl + +/-- +warning: Consider moving this declaration to the module Mathlib.Data.Nat.Notation. +note: this linter can be disabled with `set_option linter.upstreamableDecl false` +-/ +#guard_msgs in +def propose_to_move_this_def : ℕ := 0 + +-- This theorem depends on a local definition, so should not be moved. +#guard_msgs in +theorem theorem_with_local_def : propose_to_move_this_def = 0 := rfl + +-- This definition depends on definitions in two different files, so should not be moved. +#guard_msgs in +def def_with_multiple_dependencies := + let _ := Mathlib.Meta.FunProp.funPropAttr + let _ := Mathlib.Meta.NormNum.evalNatDvd + false + +end Linter.UpstreamableDecl From bcef7c52e813ca9512e907874ef563520147faf2 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 18 Nov 2024 11:19:59 +0000 Subject: [PATCH 061/829] feat(KrullDimension): coheight API (#19148) Part of #15524. This adds basic definitions for the `coheight`, and then analogous lemmas for every lemma about `height`. Due to the way `coheight` is defined, most lemmas are defeq to the corresponding lemma bout `height` in the dual order. From the Carleson project. --- Mathlib/Order/KrullDimension.lean | 143 +++++++++++++++++++++++++++++- 1 file changed, 139 insertions(+), 4 deletions(-) diff --git a/Mathlib/Order/KrullDimension.lean b/Mathlib/Order/KrullDimension.lean index 7750f256b126e..691b858cf4c86 100644 --- a/Mathlib/Order/KrullDimension.lean +++ b/Mathlib/Order/KrullDimension.lean @@ -66,7 +66,8 @@ The **coheight** of an element `a` in a preorder `α` is the supremum of the rig relation series of `α` ordered by `<` and beginning with `a`. The definition of `coheight` is via the `height` in the dual order, in order to easily transfer -theorems between `height` and `coheight`. +theorems between `height` and `coheight`. See `coheight_eq` for the definition with a +series ordered by `<` and beginning with `a`. -/ noncomputable def coheight {α : Type*} [Preorder α] (a : α) : ℕ∞ := height (α := αᵒᵈ) a @@ -82,10 +83,32 @@ variable {α β : Type*} variable [Preorder α] [Preorder β] +@[simp] lemma height_toDual (x : α) : height (OrderDual.toDual x) = coheight x := rfl +@[simp] lemma height_ofDual (x : αᵒᵈ) : height (OrderDual.ofDual x) = coheight x := rfl +@[simp] lemma coheight_toDual (x : α) : coheight (OrderDual.toDual x) = height x := rfl +@[simp] lemma coheight_ofDual (x : αᵒᵈ) : coheight (OrderDual.ofDual x) = height x := rfl + +/-- +The **coheight** of an element `a` in a preorder `α` is the supremum of the rightmost index of all +relation series of `α` ordered by `<` and beginning with `a`. + +This is not the definition of `coheight`. The definition of `coheight` is via the `height` in the +dual order, in order to easily transfer theorems between `height` and `coheight`. +-/ +lemma coheight_eq (a : α) : + coheight a = ⨆ (p : LTSeries α) (_ : a ≤ p.head), (p.length : ℕ∞) := by + apply Equiv.iSup_congr ⟨RelSeries.reverse, RelSeries.reverse, RelSeries.reverse_reverse, + RelSeries.reverse_reverse⟩ + congr! 1 + lemma height_le_iff {a : α} {n : ℕ∞} : height a ≤ n ↔ ∀ ⦃p : LTSeries α⦄, p.last ≤ a → p.length ≤ n := by rw [height, iSup₂_le_iff] +lemma coheight_le_iff {a : α} {n : ℕ∞} : + coheight a ≤ n ↔ ∀ ⦃p : LTSeries α⦄, a ≤ p.head → p.length ≤ n := by + rw [coheight_eq, iSup₂_le_iff] + lemma height_le {a : α} {n : ℕ∞} (h : ∀ (p : LTSeries α), p.last = a → p.length ≤ n) : height a ≤ n := by apply height_le_iff.mpr @@ -117,6 +140,29 @@ lemma height_eq_iSup_last_eq (a : α) : intro n rw [height_le_iff', iSup₂_le_iff] +/-- +Alternative definition of coheight, with the supremum only ranging over those series +that begin at `a`. +-/ +lemma coheight_eq_iSup_head_eq (a : α) : + coheight a = ⨆ (p : LTSeries α) (_ : p.head = a), ↑(p.length) := by + show height (α := αᵒᵈ) a = ⨆ (p : LTSeries α) (_ : p.head = a), ↑(p.length) + rw [height_eq_iSup_last_eq] + apply Equiv.iSup_congr ⟨RelSeries.reverse, RelSeries.reverse, RelSeries.reverse_reverse, + RelSeries.reverse_reverse⟩ + simp + +/-- +Variant of `coheight_le_iff` ranging only over those series that begin exactly on `a`. +-/ +lemma coheight_le_iff' {a : α} {n : ℕ∞} : + coheight a ≤ n ↔ ∀ ⦃p : LTSeries α⦄, p.head = a → p.length ≤ n := by + rw [coheight_eq_iSup_head_eq, iSup₂_le_iff] + +lemma coheight_le {a : α} {n : ℕ∞} (h : ∀ (p : LTSeries α), p.head = a → p.length ≤ n) : + coheight a ≤ n := + coheight_le_iff'.mpr h + lemma length_le_height {p : LTSeries α} {x : α} (hlast : p.last ≤ x) : p.length ≤ height x := by by_cases hlen0 : p.length ≠ 0 @@ -135,18 +181,34 @@ lemma length_le_height {p : LTSeries α} {x : α} (hlast : p.last ≤ x) : simp [p'] · simp_all +lemma length_le_coheight {x : α} {p : LTSeries α} (hhead : x ≤ p.head) : + p.length ≤ coheight x := + length_le_height (α := αᵒᵈ) (p := p.reverse) (by simpa) + /-- The height of the last element in a series is larger or equal to the length of the series. -/ lemma length_le_height_last {p : LTSeries α} : p.length ≤ height p.last := length_le_height le_rfl +/-- +The coheight of the first element in a series is larger or equal to the length of the series. +-/ +lemma length_le_coheight_head {p : LTSeries α} : p.length ≤ coheight p.head := + length_le_coheight le_rfl + /-- The height of an element in a series is larger or equal to its index in the series. -/ lemma index_le_height (p : LTSeries α) (i : Fin (p.length + 1)) : i ≤ height (p i) := length_le_height_last (p := p.take i) +/-- +The coheight of an element in a series is larger or equal to its reverse index in the series. +-/ +lemma rev_index_le_coheight (p : LTSeries α) (i : Fin (p.length + 1)) : i.rev ≤ coheight (p i) := by + simpa using index_le_height (α := αᵒᵈ) p.reverse i.rev + /-- In a maximally long series, i.e one as long as the height of the last element, the height of each element is its index in the series. @@ -160,12 +222,23 @@ lemma height_eq_index_of_length_eq_height_last {p : LTSeries α} (h : p.length = norm_cast at * omega +/-- +In a maximally long series, i.e one as long as the coheight of the first element, the coheight of +each element is its reverse index in the series. +-/ +lemma coheight_eq_index_of_length_eq_head_coheight {p : LTSeries α} (h : p.length = coheight p.head) + (i : Fin (p.length + 1)) : coheight (p i) = i.rev := by + simpa using height_eq_index_of_length_eq_height_last (α := αᵒᵈ) (p := p.reverse) (by simpa) i.rev + lemma height_mono : Monotone (α := α) height := fun _ _ hab ↦ biSup_mono (fun _ hla => hla.trans hab) @[gcongr] protected lemma _root_.GCongr.height_le_height (a b : α) (hab : a ≤ b) : height a ≤ height b := height_mono hab +lemma coheight_anti : Antitone (α := α) coheight := + (height_mono (α := αᵒᵈ)).dual_left + private lemma height_add_const (a : α) (n : ℕ∞) : height a + n = ⨆ (p : LTSeries α) (_ : p.last = a), p.length + n := by have hne : Nonempty { p : LTSeries α // p.last = a } := ⟨RelSeries.singleton _ a, rfl⟩ @@ -186,12 +259,20 @@ lemma height_le_height_apply_of_strictMono (f : α → β) (hf : StrictMono f) ( intro p hlast apply le_iSup₂_of_le (p.map f hf) (by simp [hlast]) (by simp) +lemma coheight_le_coheight_apply_of_strictMono (f : α → β) (hf : StrictMono f) (x : α) : + coheight x ≤ coheight (f x) := by + apply height_le_height_apply_of_strictMono (α := αᵒᵈ) + exact fun _ _ h ↦ hf h + @[simp] lemma height_orderIso (f : α ≃o β) (x : α) : height (f x) = height x := by apply le_antisymm · simpa using height_le_height_apply_of_strictMono _ f.symm.strictMono (f x) · exact height_le_height_apply_of_strictMono _ f.strictMono x +lemma coheight_orderIso (f : α ≃o β) (x : α) : coheight (f x) = coheight x := + height_orderIso (α := αᵒᵈ) f.dual x + private lemma exists_eq_iSup_of_iSup_eq_coe {α : Type*} [Nonempty α] {f : α → ℕ∞} {n : ℕ} (h : (⨆ x, f x) = n) : ∃ x, f x = n := by obtain ⟨x, hx⟩ := ENat.sSup_mem_of_nonempty_of_lt_top (h ▸ ENat.coe_lt_top _) @@ -222,11 +303,20 @@ lemma exists_series_of_le_height (a : α) {n : ℕ} (h : n ≤ height a) : · simp [hlast] · simp [hlen]; omega +lemma exists_series_of_le_coheight (a : α) {n : ℕ} (h : n ≤ coheight a) : + ∃ p : LTSeries α, p.head = a ∧ p.length = n := by + obtain ⟨p, hp, hl⟩ := exists_series_of_le_height (α := αᵒᵈ) a h + exact ⟨p.reverse, by simpa, by simpa⟩ + /-- For an element of finite height there exists a series ending in that element of that height. -/ lemma exists_series_of_height_eq_coe (a : α) {n : ℕ} (h : height a = n) : ∃ p : LTSeries α, p.last = a ∧ p.length = n := exists_series_of_le_height a (le_of_eq h.symm) +lemma exists_series_of_coheight_eq_coe (a : α) {n : ℕ} (h : coheight a = n) : + ∃ p : LTSeries α, p.head = a ∧ p.length = n := + exists_series_of_le_coheight a (le_of_eq h.symm) + /-- Another characterization of height, based on the supremum of the heights of elements below. -/ lemma height_eq_iSup_lt_height (x : α) : height x = ⨆ y < x, height y + 1 := by apply le_antisymm @@ -244,14 +334,22 @@ lemma height_eq_iSup_lt_height (x : α) : height x = ⨆ y < x, height y + 1 := apply iSup₂_le; intro p hp apply le_iSup₂_of_le (p.snoc x (hp ▸ hyx)) (by simp) (by simp) -lemma height_le_coe_iff {x : α} {n : ℕ} : - height x ≤ n ↔ ∀ y < x, height y < n := by +/-- +Another characterization of coheight, based on the supremum of the coheights of elements above. +-/ +lemma coheight_eq_iSup_gt_coheight (x : α) : coheight x = ⨆ y > x, coheight y + 1 := + height_eq_iSup_lt_height (α := αᵒᵈ) x + +lemma height_le_coe_iff {x : α} {n : ℕ} : height x ≤ n ↔ ∀ y < x, height y < n := by conv_lhs => rw [height_eq_iSup_lt_height, iSup₂_le_iff] congr! 2 with y _ cases height y · simp · norm_cast +lemma coheight_le_coe_iff {x : α} {n : ℕ} : coheight x ≤ n ↔ ∀ y > x, coheight y < n := + height_le_coe_iff (α := αᵒᵈ) + /-- The height of an element is infinite iff there exist series of arbitrary length ending in that element. @@ -268,15 +366,33 @@ lemma height_eq_top_iff {x : α} : obtain ⟨p, hlast, hp⟩ := h (n+1) exact ⟨p.length, ⟨⟨⟨p, hlast⟩, by simp [hp]⟩, by simp [hp]⟩⟩ +/-- +The coheight of an element is infinite iff there exist series of arbitrary length ending in that +element. +-/ +lemma coheight_eq_top_iff {x : α} : + coheight x = ⊤ ↔ ∀ n, ∃ p : LTSeries α, p.head = x ∧ p.length = n := by + convert height_eq_top_iff (α := αᵒᵈ) (x := x) using 2 with n + constructor <;> (intro ⟨p, hp, hl⟩; use p.reverse; constructor <;> simpa) + +/-- The elements of height zero are the minimal elements. -/ @[simp] lemma height_eq_zero {x : α} : height x = 0 ↔ IsMin x := by simpa [isMin_iff_forall_not_lt] using height_le_coe_iff (x := x) (n := 0) protected alias ⟨_, IsMin.height_eq_zero⟩ := height_eq_zero +/-- The elements of coheight zero are the maximal elements. -/ +@[simp] lemma coheight_eq_zero {x : α} : coheight x = 0 ↔ IsMax x := + height_eq_zero (α := αᵒᵈ) + +protected alias ⟨_, IsMax.coheight_eq_zero⟩ := coheight_eq_zero + @[simp] lemma height_bot (α : Type*) [Preorder α] [OrderBot α] : height (⊥ : α) = 0 := by simp +@[simp] lemma coheight_top (α : Type*) [Preorder α] [OrderTop α] : coheight (⊤ : α) = 0 := by simp + lemma coe_lt_height_iff {x : α} {n : ℕ} (hfin : height x < ⊤) : - n < height x ↔ (∃ y < x, height y = n) where + n < height x ↔ ∃ y < x, height y = n where mp h := by obtain ⟨m, hx : height x = m⟩ := Option.ne_none_iff_exists'.mp hfin.ne_top rw [hx] at h; norm_cast at h @@ -290,6 +406,10 @@ lemma coe_lt_height_iff {x : α} {n : ℕ} (hfin : height x < ⊤) : mpr := fun ⟨y, hyx, hy⟩ => hy ▸ height_strictMono hyx (lt_of_le_of_lt (height_mono hyx.le) hfin) +lemma coe_lt_coheight_iff {x : α} {n : ℕ} (hfin : coheight x < ⊤): + n < coheight x ↔ ∃ y > x, coheight y = n := + coe_lt_height_iff (α := αᵒᵈ) hfin + lemma height_eq_coe_add_one_iff {x : α} {n : ℕ} : height x = n + 1 ↔ height x < ⊤ ∧ (∃ y < x, height y = n) ∧ (∀ y < x, height y ≤ n) := by wlog hfin : height x < ⊤ @@ -303,6 +423,11 @@ lemma height_eq_coe_add_one_iff {x : α} {n : ℕ} : · exact coe_lt_height_iff hfin · simpa [hfin, ENat.lt_add_one_iff] using height_le_coe_iff (x := x) (n := n+1) +lemma coheight_eq_coe_add_one_iff {x : α} {n : ℕ} : + coheight x = n + 1 ↔ + coheight x < ⊤ ∧ (∃ y > x, coheight y = n) ∧ (∀ y > x, coheight y ≤ n) := + height_eq_coe_add_one_iff (α := αᵒᵈ) + lemma height_eq_coe_iff {x : α} {n : ℕ} : height x = n ↔ height x < ⊤ ∧ (n = 0 ∨ ∃ y < x, height y = n - 1) ∧ (∀ y < x, height y < n) := by @@ -319,6 +444,11 @@ lemma height_eq_coe_iff {x : α} {n : ℕ} : rename_i y _ cases height y <;> simp; norm_cast; omega +lemma coheight_eq_coe_iff {x : α} {n : ℕ} : + coheight x = n ↔ + coheight x < ⊤ ∧ (n = 0 ∨ ∃ y > x, coheight y = n - 1) ∧ (∀ y > x, coheight y < n) := + height_eq_coe_iff (α := αᵒᵈ) + /-- The elements of finite height `n` are the minimial elements among those of height `≥ n`. -/ lemma height_eq_coe_iff_minimal_le_height {a : α} {n : ℕ} : height a = n ↔ Minimal (fun y => n ≤ height y) a := by @@ -334,6 +464,11 @@ lemma height_eq_coe_iff_minimal_le_height {a : α} {n : ℕ} : use p.eraseLast.last, RelSeries.eraseLast_last_rel_last _ (by omega) simpa [hp] using length_le_height_last (p := p.eraseLast) +/-- The elements of finite coheight `n` are the maximal elements among those of coheight `≥ n`. -/ +lemma coheight_eq_coe_iff_maximal_le_coheight {a : α} {n : ℕ} : + coheight a = n ↔ Maximal (fun y => n ≤ coheight y) a := + height_eq_coe_iff_minimal_le_height (α := αᵒᵈ) + end height /-! From f4f2d31fd7b56f50b00288d83a420cfa0681f2de Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 18 Nov 2024 12:07:16 +0000 Subject: [PATCH 062/829] fix: escape `"` in `update_dependencies_zulip` (#19183) This probably fixes the action that posts reports on updating dependencies to Zulip. --- .github/workflows/update_dependencies_zulip.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.github/workflows/update_dependencies_zulip.yml b/.github/workflows/update_dependencies_zulip.yml index f4e72b01d7034..9f7270d9ee8a5 100644 --- a/.github/workflows/update_dependencies_zulip.yml +++ b/.github/workflows/update_dependencies_zulip.yml @@ -40,7 +40,7 @@ jobs: }); } } else { - output += "No PR found for this run! If you are feeling impatient and have write access, please go to the following page and click the "Run workflow" button!\nhttps://github.com/leanprover-community/mathlib4/actions/workflows/update_dependencies.yml"; + output += "No PR found for this run! If you are feeling impatient and have write access, please go to the following page and click the \"Run workflow\" button!\nhttps://github.com/leanprover-community/mathlib4/actions/workflows/update_dependencies.yml"; } return output; From 54d12203eb2afd003111899bd0a35c3ee74bb6af Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 18 Nov 2024 12:07:17 +0000 Subject: [PATCH 063/829] chore: split Topology.ContinuousMap.Bounded (#19187) --- Counterexamples/Phillips.lean | 2 +- Mathlib.lean | 3 +- Mathlib/Analysis/Normed/Lp/LpEquiv.lean | 2 +- .../Integral/RieszMarkovKakutani.lean | 2 +- .../{Bounded.lean => Bounded/Basic.lean} | 71 +------------- .../Topology/ContinuousMap/Bounded/Star.lean | 92 +++++++++++++++++++ .../BoundedCompactlySupported.lean | 2 +- Mathlib/Topology/ContinuousMap/Compact.lean | 2 +- .../Topology/ContinuousMap/ZeroAtInfty.lean | 2 +- .../MetricSpace/GromovHausdorffRealized.lean | 2 +- .../MetricSpace/ThickenedIndicator.lean | 2 +- Mathlib/Topology/Metrizable/Urysohn.lean | 2 +- Mathlib/Topology/UrysohnsBounded.lean | 2 +- 13 files changed, 106 insertions(+), 80 deletions(-) rename Mathlib/Topology/ContinuousMap/{Bounded.lean => Bounded/Basic.lean} (96%) create mode 100644 Mathlib/Topology/ContinuousMap/Bounded/Star.lean diff --git a/Counterexamples/Phillips.lean b/Counterexamples/Phillips.lean index 177816e1a065b..9b2792ab43644 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.ContinuousMap.Bounded +import Mathlib.Topology.ContinuousMap.Bounded.Star /-! # A counterexample on Pettis integrability diff --git a/Mathlib.lean b/Mathlib.lean index d1d7a4d8d578f..8aba1c993e563 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4988,7 +4988,8 @@ import Mathlib.Topology.Connected.TotallyDisconnected import Mathlib.Topology.Constructions import Mathlib.Topology.ContinuousMap.Algebra import Mathlib.Topology.ContinuousMap.Basic -import Mathlib.Topology.ContinuousMap.Bounded +import Mathlib.Topology.ContinuousMap.Bounded.Basic +import Mathlib.Topology.ContinuousMap.Bounded.Star import Mathlib.Topology.ContinuousMap.BoundedCompactlySupported import Mathlib.Topology.ContinuousMap.CocompactMap import Mathlib.Topology.ContinuousMap.Compact diff --git a/Mathlib/Analysis/Normed/Lp/LpEquiv.lean b/Mathlib/Analysis/Normed/Lp/LpEquiv.lean index 013c08de891cd..6635a5d5fb355 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.ContinuousMap.Bounded +import Mathlib.Topology.ContinuousMap.Bounded.Basic /-! # Equivalences among $L^p$ spaces diff --git a/Mathlib/MeasureTheory/Integral/RieszMarkovKakutani.lean b/Mathlib/MeasureTheory/Integral/RieszMarkovKakutani.lean index da3bb92699829..0c8f7657a9261 100644 --- a/Mathlib/MeasureTheory/Integral/RieszMarkovKakutani.lean +++ b/Mathlib/MeasureTheory/Integral/RieszMarkovKakutani.lean @@ -3,8 +3,8 @@ 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.ContinuousMap.Bounded import Mathlib.Topology.Sets.Compacts +import Mathlib.Topology.ContinuousMap.Bounded.Basic /-! # Riesz–Markov–Kakutani representation theorem diff --git a/Mathlib/Topology/ContinuousMap/Bounded.lean b/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean similarity index 96% rename from Mathlib/Topology/ContinuousMap/Bounded.lean rename to Mathlib/Topology/ContinuousMap/Bounded/Basic.lean index 0c622b522b9ad..40d279aec7e3a 100644 --- a/Mathlib/Topology/ContinuousMap/Bounded.lean +++ b/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean @@ -7,7 +7,6 @@ import Mathlib.Algebra.Module.MinimalAxioms import Mathlib.Topology.ContinuousMap.Algebra 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 @@ -19,6 +18,8 @@ the uniform distance. -/ +assert_not_exists CStarAlgebra.star + noncomputable section open Topology Bornology NNReal uniformity UniformConvergence @@ -1367,72 +1368,6 @@ theorem NNReal.upper_bound {α : Type*} [TopologicalSpace α] (f : α →ᵇ ℝ simp only [coe_zero, Pi.zero_apply] at key rwa [NNReal.nndist_zero_eq_val' (f x)] at key -/-! -### Star structures - -In this section, if `β` is a normed ⋆-group, then so is the space of bounded -continuous functions from `α` to `β`, by using the star operation pointwise. - -If `𝕜` is normed field and a ⋆-ring over which `β` is a normed algebra and a -star module, then the space of bounded continuous functions from `α` to `β` -is a star module. - -If `β` is a ⋆-ring in addition to being a normed ⋆-group, then `α →ᵇ β` -inherits a ⋆-ring structure. - -In summary, if `β` is a C⋆-algebra over `𝕜`, then so is `α →ᵇ β`; note that -completeness is guaranteed when `β` is complete (see -`BoundedContinuousFunction.complete`). -/ - - -section NormedAddCommGroup - -variable {𝕜 : Type*} [NormedField 𝕜] [StarRing 𝕜] [TopologicalSpace α] [SeminormedAddCommGroup β] - [StarAddMonoid β] [NormedStarGroup β] - -variable [NormedSpace 𝕜 β] [StarModule 𝕜 β] - -instance instStarAddMonoid : StarAddMonoid (α →ᵇ β) where - star f := f.comp star starNormedAddGroupHom.lipschitz - star_involutive f := ext fun x => star_star (f x) - star_add f g := ext fun x => star_add (f x) (g x) - -/-- The right-hand side of this equality can be parsed `star ∘ ⇑f` because of the -instance `Pi.instStarForAll`. Upon inspecting the goal, one sees `⊢ ↑(star f) = star ↑f`. -/ -@[simp] -theorem coe_star (f : α →ᵇ β) : ⇑(star f) = star (⇑f) := rfl - -@[simp] -theorem star_apply (f : α →ᵇ β) (x : α) : star f x = star (f x) := rfl - -instance instNormedStarGroup : NormedStarGroup (α →ᵇ β) where - norm_star f := by simp only [norm_eq, star_apply, norm_star] - -instance instStarModule : StarModule 𝕜 (α →ᵇ β) where - star_smul k f := ext fun x => star_smul k (f x) - -end NormedAddCommGroup - -section CStarRing - -variable [TopologicalSpace α] -variable [NonUnitalNormedRing β] [StarRing β] - -instance instStarRing [NormedStarGroup β] : StarRing (α →ᵇ β) where - __ := instStarAddMonoid - star_mul f g := ext fun x ↦ star_mul (f x) (g x) - -variable [CStarRing β] - -instance instCStarRing : CStarRing (α →ᵇ β) where - norm_mul_self_le f := by - rw [← sq, ← Real.le_sqrt (norm_nonneg _) (norm_nonneg _), norm_le (Real.sqrt_nonneg _)] - intro x - rw [Real.le_sqrt (norm_nonneg _) (norm_nonneg _), sq, ← CStarRing.norm_star_mul_self] - exact norm_coe_le_norm (star f * f) x - -end CStarRing - section NormedLatticeOrderedGroup variable [TopologicalSpace α] [NormedLatticeAddCommGroup β] @@ -1560,5 +1495,3 @@ lemma norm_sub_nonneg (f : α →ᵇ ℝ) : end end BoundedContinuousFunction - -set_option linter.style.longFile 1700 diff --git a/Mathlib/Topology/ContinuousMap/Bounded/Star.lean b/Mathlib/Topology/ContinuousMap/Bounded/Star.lean new file mode 100644 index 0000000000000..c89449232bdaa --- /dev/null +++ b/Mathlib/Topology/ContinuousMap/Bounded/Star.lean @@ -0,0 +1,92 @@ +/- +Copyright (c) 2018 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel, Mario Carneiro, Yury Kudryashov, Heather Macbeth +-/ +import Mathlib.Topology.ContinuousMap.Bounded.Basic +import Mathlib.Analysis.CStarAlgebra.Basic + +/-! +# Star structures on bounded continuous functions + +-/ + +noncomputable section + +open Topology Bornology NNReal uniformity UniformConvergence + +open Set Filter Metric Function + +universe u v w + +variable {F : Type*} {α : Type u} {β : Type v} {γ : Type w} + +namespace BoundedContinuousFunction + +/-! +### Star structures + +In this section, if `β` is a normed ⋆-group, then so is the space of bounded +continuous functions from `α` to `β`, by using the star operation pointwise. + +If `𝕜` is normed field and a ⋆-ring over which `β` is a normed algebra and a +star module, then the space of bounded continuous functions from `α` to `β` +is a star module. + +If `β` is a ⋆-ring in addition to being a normed ⋆-group, then `α →ᵇ β` +inherits a ⋆-ring structure. + +In summary, if `β` is a C⋆-algebra over `𝕜`, then so is `α →ᵇ β`; note that +completeness is guaranteed when `β` is complete (see +`BoundedContinuousFunction.complete`). -/ + + +section NormedAddCommGroup + +variable {𝕜 : Type*} [NormedField 𝕜] [StarRing 𝕜] [TopologicalSpace α] [SeminormedAddCommGroup β] + [StarAddMonoid β] [NormedStarGroup β] + +variable [NormedSpace 𝕜 β] [StarModule 𝕜 β] + +instance instStarAddMonoid : StarAddMonoid (α →ᵇ β) where + star f := f.comp star starNormedAddGroupHom.lipschitz + star_involutive f := ext fun x => star_star (f x) + star_add f g := ext fun x => star_add (f x) (g x) + +/-- The right-hand side of this equality can be parsed `star ∘ ⇑f` because of the +instance `Pi.instStarForAll`. Upon inspecting the goal, one sees `⊢ ↑(star f) = star ↑f`. -/ +@[simp] +theorem coe_star (f : α →ᵇ β) : ⇑(star f) = star (⇑f) := rfl + +@[simp] +theorem star_apply (f : α →ᵇ β) (x : α) : star f x = star (f x) := rfl + +instance instNormedStarGroup : NormedStarGroup (α →ᵇ β) where + norm_star f := by simp only [norm_eq, star_apply, norm_star] + +instance instStarModule : StarModule 𝕜 (α →ᵇ β) where + star_smul k f := ext fun x => star_smul k (f x) + +end NormedAddCommGroup + +section CStarRing + +variable [TopologicalSpace α] +variable [NonUnitalNormedRing β] [StarRing β] + +instance instStarRing [NormedStarGroup β] : StarRing (α →ᵇ β) where + __ := instStarAddMonoid + star_mul f g := ext fun x ↦ star_mul (f x) (g x) + +variable [CStarRing β] + +instance instCStarRing : CStarRing (α →ᵇ β) where + norm_mul_self_le f := by + rw [← sq, ← Real.le_sqrt (norm_nonneg _) (norm_nonneg _), norm_le (Real.sqrt_nonneg _)] + intro x + rw [Real.le_sqrt (norm_nonneg _) (norm_nonneg _), sq, ← CStarRing.norm_star_mul_self] + exact norm_coe_le_norm (star f * f) x + +end CStarRing + +end BoundedContinuousFunction diff --git a/Mathlib/Topology/ContinuousMap/BoundedCompactlySupported.lean b/Mathlib/Topology/ContinuousMap/BoundedCompactlySupported.lean index ff7efe16ef944..6e4eac4a1bcd9 100644 --- a/Mathlib/Topology/ContinuousMap/BoundedCompactlySupported.lean +++ b/Mathlib/Topology/ContinuousMap/BoundedCompactlySupported.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.ContinuousMap.Bounded import Mathlib.RingTheory.TwoSidedIdeal.Lattice +import Mathlib.Topology.ContinuousMap.Bounded.Basic /-! # Compactly supported bounded continuous functions diff --git a/Mathlib/Topology/ContinuousMap/Compact.lean b/Mathlib/Topology/ContinuousMap/Compact.lean index 7c7444817b6fc..e6cb2f048e863 100644 --- a/Mathlib/Topology/ContinuousMap/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.ContinuousMap.Bounded +import Mathlib.Topology.ContinuousMap.Bounded.Star import Mathlib.Topology.ContinuousMap.Star import Mathlib.Topology.UniformSpace.Compact import Mathlib.Topology.CompactOpen diff --git a/Mathlib/Topology/ContinuousMap/ZeroAtInfty.lean b/Mathlib/Topology/ContinuousMap/ZeroAtInfty.lean index 2f8e5d908c1a5..eecb451c02ae7 100644 --- a/Mathlib/Topology/ContinuousMap/ZeroAtInfty.lean +++ b/Mathlib/Topology/ContinuousMap/ZeroAtInfty.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.ContinuousMap.Bounded +import Mathlib.Topology.ContinuousMap.Bounded.Star import Mathlib.Topology.ContinuousMap.CocompactMap /-! diff --git a/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean b/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean index 64b4f2c7639bf..c7494ebd7c280 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.ContinuousMap.Bounded +import Mathlib.Topology.ContinuousMap.Bounded.Basic /-! # The Gromov-Hausdorff distance is realized diff --git a/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean b/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean index b3e9cec910f90..b30ad210dc644 100644 --- a/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean +++ b/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kalle Kytölä -/ import Mathlib.Data.ENNReal.Lemmas -import Mathlib.Topology.ContinuousMap.Bounded import Mathlib.Topology.MetricSpace.Thickening +import Mathlib.Topology.ContinuousMap.Bounded.Basic /-! # Thickened indicators diff --git a/Mathlib/Topology/Metrizable/Urysohn.lean b/Mathlib/Topology/Metrizable/Urysohn.lean index 654118a44100c..9d41da4d5bc28 100644 --- a/Mathlib/Topology/Metrizable/Urysohn.lean +++ b/Mathlib/Topology/Metrizable/Urysohn.lean @@ -5,8 +5,8 @@ Authors: Yury Kudryashov -/ import Mathlib.Analysis.SpecificLimits.Basic import Mathlib.Topology.UrysohnsLemma -import Mathlib.Topology.ContinuousMap.Bounded import Mathlib.Topology.Metrizable.Basic +import Mathlib.Topology.ContinuousMap.Bounded.Basic /-! # Urysohn's Metrization Theorem diff --git a/Mathlib/Topology/UrysohnsBounded.lean b/Mathlib/Topology/UrysohnsBounded.lean index 16835df30f181..5be5913ea5318 100644 --- a/Mathlib/Topology/UrysohnsBounded.lean +++ b/Mathlib/Topology/UrysohnsBounded.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Topology.UrysohnsLemma -import Mathlib.Topology.ContinuousMap.Bounded +import Mathlib.Topology.ContinuousMap.Bounded.Basic /-! # Urysohn's lemma for bounded continuous functions From 5fbd4a73e40dc0c21f722c1b1671285ea75e8f21 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 18 Nov 2024 12:21:49 +0000 Subject: [PATCH 064/829] chore: move `Fin` material earlier (#19186) This will in particular be used in #19086 to prove `Finite Bool` and `Finite Prop` much earlier. --- Mathlib/Algebra/Ring/Fin.lean | 2 +- .../FiniteProductsOfBinaryProducts.lean | 1 - .../CategoryTheory/Monoidal/Types/Basic.lean | 1 - .../SimpleGraph/ConcreteColorings.lean | 1 - Mathlib/Data/Countable/Basic.lean | 4 +- Mathlib/Data/Fin/Basic.lean | 3 - Mathlib/Data/Fin/Tuple/Basic.lean | 52 ++++++++++++ Mathlib/Data/Fin/Tuple/Finset.lean | 1 - Mathlib/Data/ZMod/Defs.lean | 2 +- Mathlib/Logic/Equiv/Defs.lean | 19 +++++ Mathlib/Logic/Equiv/Fin.lean | 81 ------------------- Mathlib/Logic/Unique.lean | 2 + .../MeasureTheory/MeasurableSpace/Basic.lean | 1 - Mathlib/SetTheory/Cardinal/Basic.lean | 1 + 14 files changed, 78 insertions(+), 93 deletions(-) diff --git a/Mathlib/Algebra/Ring/Fin.lean b/Mathlib/Algebra/Ring/Fin.lean index 6f993f031cb5d..d3e75cc19cd0a 100644 --- a/Mathlib/Algebra/Ring/Fin.lean +++ b/Mathlib/Algebra/Ring/Fin.lean @@ -5,7 +5,7 @@ Authors: Anne Baanen -/ import Mathlib.Algebra.Group.Prod import Mathlib.Algebra.Ring.Equiv -import Mathlib.Logic.Equiv.Fin +import Mathlib.Data.Fin.Tuple.Basic /-! # Rings and `Fin` diff --git a/Mathlib/CategoryTheory/Limits/Constructions/FiniteProductsOfBinaryProducts.lean b/Mathlib/CategoryTheory/Limits/Constructions/FiniteProductsOfBinaryProducts.lean index 83fd7abbab1b6..ffb743849fa51 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/FiniteProductsOfBinaryProducts.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/FiniteProductsOfBinaryProducts.lean @@ -7,7 +7,6 @@ import Mathlib.CategoryTheory.Limits.Preserves.Shapes.BinaryProducts import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Products import Mathlib.CategoryTheory.Limits.Shapes.BinaryProducts import Mathlib.CategoryTheory.Limits.Shapes.FiniteProducts -import Mathlib.Logic.Equiv.Fin /-! # Constructing finite products from binary products and terminal. diff --git a/Mathlib/CategoryTheory/Monoidal/Types/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Types/Basic.lean index 155636639d804..f0c019b44cef8 100644 --- a/Mathlib/CategoryTheory/Monoidal/Types/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Types/Basic.lean @@ -6,7 +6,6 @@ Authors: Michael Jendrusch, Kim Morrison import Mathlib.CategoryTheory.Monoidal.Functor import Mathlib.CategoryTheory.ChosenFiniteProducts import Mathlib.CategoryTheory.Limits.Shapes.Types -import Mathlib.Logic.Equiv.Fin /-! # The category of types is a monoidal category diff --git a/Mathlib/Combinatorics/SimpleGraph/ConcreteColorings.lean b/Mathlib/Combinatorics/SimpleGraph/ConcreteColorings.lean index 33ece69cf7414..b968d04daf2bc 100644 --- a/Mathlib/Combinatorics/SimpleGraph/ConcreteColorings.lean +++ b/Mathlib/Combinatorics/SimpleGraph/ConcreteColorings.lean @@ -5,7 +5,6 @@ Authors: Iván Renison -/ import Mathlib.Combinatorics.SimpleGraph.Coloring import Mathlib.Combinatorics.SimpleGraph.Hasse -import Mathlib.Logic.Equiv.Fin /-! # Concrete colorings of common graphs diff --git a/Mathlib/Data/Countable/Basic.lean b/Mathlib/Data/Countable/Basic.lean index ccd6697716ddf..caecc1a3b4c19 100644 --- a/Mathlib/Data/Countable/Basic.lean +++ b/Mathlib/Data/Countable/Basic.lean @@ -3,9 +3,9 @@ Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import Mathlib.Logic.Equiv.Nat -import Mathlib.Logic.Equiv.Fin import Mathlib.Data.Countable.Defs +import Mathlib.Data.Fin.Tuple.Basic +import Mathlib.Logic.Equiv.Nat /-! # Countable types diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 017d7dc2c1409..21c9e27353985 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -1465,9 +1465,6 @@ protected theorem coe_neg (a : Fin n) : ((-a : Fin n) : ℕ) = (n - a) % n := theorem eq_zero (n : Fin 1) : n = 0 := Subsingleton.elim _ _ -instance uniqueFinOne : Unique (Fin 1) where - uniq _ := Subsingleton.elim _ _ - @[deprecated val_eq_zero (since := "2024-09-18")] theorem coe_fin_one (a : Fin 1) : (a : ℕ) = 0 := by simp [Subsingleton.elim a 0] diff --git a/Mathlib/Data/Fin/Tuple/Basic.lean b/Mathlib/Data/Fin/Tuple/Basic.lean index 63c1bba055d5c..dd1645bcb7380 100644 --- a/Mathlib/Data/Fin/Tuple/Basic.lean +++ b/Mathlib/Data/Fin/Tuple/Basic.lean @@ -177,6 +177,18 @@ theorem cons_self_tail : cons (q 0) (tail q) = q := by unfold tail rw [cons_succ] +/-- Equivalence between tuples of length `n + 1` and pairs of an element and a tuple of length `n` +given by separating out the first element of the tuple. + +This is `Fin.cons` as an `Equiv`. -/ +@[simps] +def consEquiv (α : Fin (n + 1) → Type*) : α 0 × (∀ i, α (succ i)) ≃ ∀ i, α i where + toFun f := cons f.1 f.2 + invFun f := (f 0, tail f) + left_inv f := by simp + right_inv f := by simp + + -- Porting note: Mathport removes `_root_`? /-- Recurse on an `n+1`-tuple by splitting it into a single element and an `n`-tuple. -/ @[elab_as_elim] @@ -693,6 +705,17 @@ theorem comp_init {α : Sort*} {β : Sort*} (g : α → β) (q : Fin n.succ → ext j simp [init] +/-- Equivalence between tuples of length `n + 1` and pairs of an element and a tuple of length `n` +given by separating out the last element of the tuple. + +This is `Fin.snoc` as an `Equiv`. -/ +@[simps] +def snocEquiv (α : Fin (n + 1) → Type*) : α (last n) × (∀ i, α (castSucc i)) ≃ ∀ i, α i where + toFun f _ := Fin.snoc f.2 f.1 _ + invFun f := ⟨f _, Fin.init f⟩ + left_inv f := by simp + right_inv f := by simp + /-- Recurse on an `n+1`-tuple by splitting it its initial `n`-tuple and its last element. -/ @[elab_as_elim, inline] def snocCases {P : (∀ i : Fin n.succ, α i) → Sort*} @@ -926,6 +949,26 @@ open Set lemma insertNth_self_removeNth (p : Fin (n + 1)) (f : ∀ j, α j) : insertNth p (f p) (removeNth p f) = f := by simp +/-- Equivalence between tuples of length `n + 1` and pairs of an element and a tuple of length `n` +given by separating out the `p`-th element of the tuple. + +This is `Fin.insertNth` as an `Equiv`. -/ +@[simps] +def insertNthEquiv (α : Fin (n + 1) → Type u) (p : Fin (n + 1)) : + α p × (∀ i, α (p.succAbove i)) ≃ ∀ i, α i where + toFun f := insertNth p f.1 f.2 + invFun f := (f p, removeNth p f) + left_inv f := by ext <;> simp + right_inv f := by simp + +@[simp] lemma insertNthEquiv_zero (α : Fin (n + 1) → Type*) : insertNthEquiv α 0 = consEquiv α := + Equiv.symm_bijective.injective <| by ext <;> rfl + +/-- Note this lemma can only be written about non-dependent tuples as `insertNth (last n) = snoc` is +not a definitional equality. -/ +@[simp] lemma insertNthEquiv_last (n : ℕ) (α : Type*) : + insertNthEquiv (fun _ ↦ α) (last n) = snocEquiv (fun _ ↦ α) := by ext; simp + /-- Separates an `n+1`-tuple, returning a selected index and then the rest of the tuple. Functional form of `Equiv.piFinSuccAbove`. -/ @[deprecated removeNth (since := "2024-06-19")] @@ -1090,3 +1133,12 @@ theorem sigma_eq_iff_eq_comp_cast {α : Type*} {a b : Σii, Fin ii → α} : sigma_eq_of_eq_comp_cast _ h'⟩ end Fin + +/-- `Π i : Fin 2, α i` is equivalent to `α 0 × α 1`. See also `finTwoArrowEquiv` for a +non-dependent version and `prodEquivPiFinTwo` for a version with inputs `α β : Type u`. -/ +@[simps (config := .asFn)] +def piFinTwoEquiv (α : Fin 2 → Type u) : (∀ i, α i) ≃ α 0 × α 1 where + toFun f := (f 0, f 1) + invFun p := Fin.cons p.1 <| Fin.cons p.2 finZeroElim + left_inv _ := funext <| Fin.forall_fin_two.2 ⟨rfl, rfl⟩ + right_inv := fun _ => rfl diff --git a/Mathlib/Data/Fin/Tuple/Finset.lean b/Mathlib/Data/Fin/Tuple/Finset.lean index 7c012e3737704..7704dde3c87fb 100644 --- a/Mathlib/Data/Fin/Tuple/Finset.lean +++ b/Mathlib/Data/Fin/Tuple/Finset.lean @@ -5,7 +5,6 @@ Authors: Bolton Bailey -/ import Mathlib.Data.Finset.Prod import Mathlib.Data.Fintype.Pi -import Mathlib.Logic.Equiv.Fin /-! # Fin-indexed tuples of finsets diff --git a/Mathlib/Data/ZMod/Defs.lean b/Mathlib/Data/ZMod/Defs.lean index 80aa3c4ee070c..a1371c6d87fce 100644 --- a/Mathlib/Data/ZMod/Defs.lean +++ b/Mathlib/Data/ZMod/Defs.lean @@ -101,7 +101,7 @@ instance ZMod.repr : ∀ n : ℕ, Repr (ZMod n) namespace ZMod -instance instUnique : Unique (ZMod 1) := Fin.uniqueFinOne +instance instUnique : Unique (ZMod 1) := Fin.instUnique instance fintype : ∀ (n : ℕ) [NeZero n], Fintype (ZMod n) | 0, h => (h.ne _ rfl).elim diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index 2650567700550..9c482264afc23 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -854,3 +854,22 @@ protected def congrRight {r r' : Setoid α} Quot.congrRight eq end Quotient + +/-- Equivalence between `Fin 0` and `Empty`. -/ +def finZeroEquiv : Fin 0 ≃ Empty := .equivEmpty _ + +/-- Equivalence between `Fin 0` and `PEmpty`. -/ +def finZeroEquiv' : Fin 0 ≃ PEmpty.{u} := .equivPEmpty _ + +/-- Equivalence between `Fin 1` and `Unit`. -/ +def finOneEquiv : Fin 1 ≃ Unit := .equivPUnit _ + +/-- Equivalence between `Fin 2` and `Bool`. -/ +def finTwoEquiv : Fin 2 ≃ Bool where + toFun i := i == 1 + invFun b := bif b then 1 else 0 + left_inv i := + match i with + | 0 => by simp + | 1 => by simp + right_inv b := by cases b <;> simp diff --git a/Mathlib/Logic/Equiv/Fin.lean b/Mathlib/Logic/Equiv/Fin.lean index 7e551963a14c7..a8dc54a1bbf67 100644 --- a/Mathlib/Logic/Equiv/Fin.lean +++ b/Mathlib/Logic/Equiv/Fin.lean @@ -18,87 +18,6 @@ universe u variable {m n : ℕ} -/-- Equivalence between `Fin 0` and `Empty`. -/ -def finZeroEquiv : Fin 0 ≃ Empty := - Equiv.equivEmpty _ - -/-- Equivalence between `Fin 0` and `PEmpty`. -/ -def finZeroEquiv' : Fin 0 ≃ PEmpty.{u} := - Equiv.equivPEmpty _ - -/-- Equivalence between `Fin 1` and `Unit`. -/ -def finOneEquiv : Fin 1 ≃ Unit := - Equiv.equivPUnit _ - -/-- Equivalence between `Fin 2` and `Bool`. -/ -def finTwoEquiv : Fin 2 ≃ Bool where - toFun := ![false, true] - invFun b := b.casesOn 0 1 - left_inv := Fin.forall_fin_two.2 <| by simp - right_inv := Bool.forall_bool.2 <| by simp - -/-! -### Tuples - -This section defines a bunch of equivalences between `n + 1`-tuples and products of `n`-tuples with -an entry. --/ - -namespace Fin - -/-- Equivalence between tuples of length `n + 1` and pairs of an element and a tuple of length `n` -given by separating out the first element of the tuple. - -This is `Fin.cons` as an `Equiv`. -/ -@[simps] -def consEquiv (α : Fin (n + 1) → Type*) : α 0 × (∀ i, α (succ i)) ≃ ∀ i, α i where - toFun f := cons f.1 f.2 - invFun f := (f 0, tail f) - left_inv f := by simp - right_inv f := by simp - -/-- Equivalence between tuples of length `n + 1` and pairs of an element and a tuple of length `n` -given by separating out the last element of the tuple. - -This is `Fin.snoc` as an `Equiv`. -/ -@[simps] -def snocEquiv (α : Fin (n + 1) → Type*) : α (last n) × (∀ i, α (castSucc i)) ≃ ∀ i, α i where - toFun f _ := Fin.snoc f.2 f.1 _ - invFun f := ⟨f _, Fin.init f⟩ - left_inv f := by simp - right_inv f := by simp - -/-- Equivalence between tuples of length `n + 1` and pairs of an element and a tuple of length `n` -given by separating out the `p`-th element of the tuple. - -This is `Fin.insertNth` as an `Equiv`. -/ -@[simps] -def insertNthEquiv (α : Fin (n + 1) → Type u) (p : Fin (n + 1)) : - α p × (∀ i, α (p.succAbove i)) ≃ ∀ i, α i where - toFun f := insertNth p f.1 f.2 - invFun f := (f p, removeNth p f) - left_inv f := by ext <;> simp - right_inv f := by simp - -@[simp] lemma insertNthEquiv_zero (α : Fin (n + 1) → Type*) : insertNthEquiv α 0 = consEquiv α := - Equiv.symm_bijective.injective <| by ext <;> rfl - -/-- Note this lemma can only be written about non-dependent tuples as `insertNth (last n) = snoc` is -not a definitional equality. -/ -@[simp] lemma insertNthEquiv_last (n : ℕ) (α : Type*) : - insertNthEquiv (fun _ ↦ α) (last n) = snocEquiv (fun _ ↦ α) := by ext; simp - -end Fin - -/-- `Π i : Fin 2, α i` is equivalent to `α 0 × α 1`. See also `finTwoArrowEquiv` for a -non-dependent version and `prodEquivPiFinTwo` for a version with inputs `α β : Type u`. -/ -@[simps (config := .asFn)] -def piFinTwoEquiv (α : Fin 2 → Type u) : (∀ i, α i) ≃ α 0 × α 1 where - toFun f := (f 0, f 1) - invFun p := Fin.cons p.1 <| Fin.cons p.2 finZeroElim - left_inv _ := funext <| Fin.forall_fin_two.2 ⟨rfl, rfl⟩ - right_inv := fun _ => rfl - /-! ### Miscellaneous diff --git a/Mathlib/Logic/Unique.lean b/Mathlib/Logic/Unique.lean index 5aad9af23a014..5c28b89ae00cf 100644 --- a/Mathlib/Logic/Unique.lean +++ b/Mathlib/Logic/Unique.lean @@ -261,3 +261,5 @@ instance Unique.subtypeEq' (y : α) : Unique { x // y = x } where uniq := fun ⟨x, hx⟩ ↦ by subst hx; congr end Subtype + +instance Fin.instUnique : Unique (Fin 1) where uniq _ := Subsingleton.elim _ _ diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index db568c4737312..4d2eb7fd31ab1 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -8,7 +8,6 @@ import Mathlib.Data.Int.Cast.Lemmas import Mathlib.Data.Prod.TProd import Mathlib.Data.Set.UnionLift import Mathlib.GroupTheory.Coset.Defs -import Mathlib.Logic.Equiv.Fin import Mathlib.MeasureTheory.MeasurableSpace.Instances import Mathlib.Order.Filter.SmallSets import Mathlib.Tactic.FinCases diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index 40f07f1cc48f2..18f50f95db1cb 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -8,6 +8,7 @@ import Mathlib.Data.Fintype.BigOperators import Mathlib.Data.Fintype.Powerset import Mathlib.Data.Nat.Cast.Order.Basic import Mathlib.Data.Set.Countable +import Mathlib.Logic.Equiv.Fin import Mathlib.Logic.Small.Set import Mathlib.Logic.UnivLE import Mathlib.Order.ConditionallyCompleteLattice.Indexed From 7698f92f789c146c3f0a5015789963682e18cc46 Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Mon, 18 Nov 2024 13:02:28 +0000 Subject: [PATCH 065/829] feat: add IsPrimitiveRoot.pow_sub_pow_eq_prod_sub_mul (#19022) From flt-regular. --- Mathlib/Algebra/Polynomial/Roots.lean | 8 ++++ .../Polynomial/Cyclotomic/Basic.lean | 46 ++++++++++++++++++- 2 files changed, 53 insertions(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Polynomial/Roots.lean b/Mathlib/Algebra/Polynomial/Roots.lean index ac9cd6e553c25..d2a32fe42bc6e 100644 --- a/Mathlib/Algebra/Polynomial/Roots.lean +++ b/Mathlib/Algebra/Polynomial/Roots.lean @@ -312,6 +312,14 @@ theorem mem_nthRootsFinset {n : ℕ} (h : 0 < n) {x : R} : @[simp] theorem nthRootsFinset_zero : nthRootsFinset 0 R = ∅ := by classical simp [nthRootsFinset_def] +theorem map_mem_nthRootsFinset {S F : Type*} [CommRing S] [IsDomain S] [FunLike F R S] + [RingHomClass F R S] {x : R} (hx : x ∈ nthRootsFinset n R) (f : F) : + f x ∈ nthRootsFinset n S := by + by_cases hn : n = 0 + · simp [hn] at hx + · rw [mem_nthRootsFinset <| Nat.pos_of_ne_zero hn, ← map_pow, (mem_nthRootsFinset <| + Nat.pos_of_ne_zero hn).1 hx, map_one] + theorem mul_mem_nthRootsFinset {η₁ η₂ : R} (hη₁ : η₁ ∈ nthRootsFinset n R) (hη₂ : η₂ ∈ nthRootsFinset n R) : η₁ * η₂ ∈ nthRootsFinset n R := by diff --git a/Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean b/Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean index bd7bd57570e25..179f1fea4b15a 100644 --- a/Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean @@ -582,13 +582,57 @@ end Order section miscellaneous -lemma dvd_C_mul_X_sub_one_pow_add_one {R : Type*} [CommRing R] {p : ℕ} (hpri : p.Prime) +open Finset + +variable {R : Type*} [CommRing R] {ζ : R} {n : ℕ} (x y : R) + +lemma dvd_C_mul_X_sub_one_pow_add_one {p : ℕ} (hpri : p.Prime) (hp : p ≠ 2) (a r : R) (h₁ : r ∣ a ^ p) (h₂ : r ∣ p * a) : C r ∣ (C a * X - 1) ^ p + 1 := by have := hpri.dvd_add_pow_sub_pow_of_dvd (C a * X) (-1) (r := C r) ?_ ?_ · rwa [← sub_eq_add_neg, (hpri.odd_of_ne_two hp).neg_pow, one_pow, sub_neg_eq_add] at this · simp only [mul_pow, ← map_pow, dvd_mul_right, (_root_.map_dvd C h₁).trans] simp only [map_mul, map_natCast, ← mul_assoc, dvd_mul_right, (_root_.map_dvd C h₂).trans] +private theorem _root_.IsPrimitiveRoot.pow_sub_pow_eq_prod_sub_mul_field {K : Type*} + [Field K] {ζ : K} (x y : K) (hpos : 0 < n) (h : IsPrimitiveRoot ζ n) : + x ^ n - y ^ n = ∏ ζ ∈ nthRootsFinset n K, (x - ζ * y) := by + by_cases hy : y = 0 + · simp only [hy, zero_pow (Nat.not_eq_zero_of_lt hpos), sub_zero, mul_zero, prod_const] + congr + rw [h.card_nthRootsFinset] + convert congr_arg (eval (x/y) · * y ^ card (nthRootsFinset n K)) <| X_pow_sub_one_eq_prod hpos h + using 1 + · simp [sub_mul, div_pow, hy, h.card_nthRootsFinset] + · simp [eval_prod, prod_mul_pow_card, sub_mul, hy] + +variable [IsDomain R] + +/-- If there is a primitive `n`th root of unity in `R`, then `X ^ n - Y ^ n = ∏ (X - μ Y)`, +where `μ` varies over the `n`-th roots of unity. -/ +theorem _root_.IsPrimitiveRoot.pow_sub_pow_eq_prod_sub_mul (hpos : 0 < n) + (h : IsPrimitiveRoot ζ n) : x ^ n - y ^ n = ∏ ζ ∈ nthRootsFinset n R, (x - ζ * y) := by + let K := FractionRing R + apply NoZeroSMulDivisors.algebraMap_injective R K + rw [map_sub, map_pow, map_pow, map_prod] + simp_rw [map_sub, map_mul] + have h' : IsPrimitiveRoot (algebraMap R K ζ) n := + h.map_of_injective <| NoZeroSMulDivisors.algebraMap_injective R K + rw [h'.pow_sub_pow_eq_prod_sub_mul_field _ _ hpos] + refine (prod_nbij (algebraMap R K) (fun a ha ↦ map_mem_nthRootsFinset ha _) (fun a _ b _ H ↦ + NoZeroSMulDivisors.algebraMap_injective R K H) (fun a ha ↦ ?_) (fun _ _ ↦ rfl)).symm + have := Set.surj_on_of_inj_on_of_ncard_le (s := nthRootsFinset n R) + (t := nthRootsFinset n K) _ (fun _ hr ↦ map_mem_nthRootsFinset hr _) + (fun a _ b _ H ↦ NoZeroSMulDivisors.algebraMap_injective R K H) + (by simp [h.card_nthRootsFinset, h'.card_nthRootsFinset]) + obtain ⟨x, hx, hx1⟩ := this _ ha + exact ⟨x, hx, hx1.symm⟩ + +/-- If there is a primitive `n`th root of unity in `R` and `n` is odd, then +`X ^ n + Y ^ n = ∏ (X + μ Y)`, where `μ` varies over the `n`-th roots of unity. -/ +theorem _root_.IsPrimitiveRoot.pow_add_pow_eq_prod_add_mul (hodd : Odd n) + (h : IsPrimitiveRoot ζ n) : x ^ n + y ^ n = ∏ ζ ∈ nthRootsFinset n R, (x + ζ * y) := by + simpa [hodd.neg_pow] using h.pow_sub_pow_eq_prod_sub_mul x (-y) hodd.pos + end miscellaneous end Polynomial From 2c6450e45a0b2b45662a8289c3d0fbf09883939f Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Mon, 18 Nov 2024 13:17:54 +0000 Subject: [PATCH 066/829] feat(CategoryTheory): a biproduct of epimorphisms is an epimorphism, etc (#19190) --- .../Limits/Shapes/Biproducts.lean | 55 +++++++++++++++++++ 1 file changed, 55 insertions(+) diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean index e411b582615b0..dfe5ff42f7453 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean @@ -610,6 +610,37 @@ def biproduct.mapIso {f g : J → C} [HasBiproduct f] [HasBiproduct g] (p : ∀ hom := biproduct.map fun b => (p b).hom inv := biproduct.map fun b => (p b).inv +instance biproduct.map_epi {f g : J → C} [HasBiproduct f] [HasBiproduct g] (p : ∀ j, f j ⟶ g j) + [∀ j, Epi (p j)] : Epi (biproduct.map p) := by + have : biproduct.map p = + (biproduct.isoCoproduct _).hom ≫ Sigma.map p ≫ (biproduct.isoCoproduct _).inv := by + ext + simp only [map_π, isoCoproduct_hom, isoCoproduct_inv, Category.assoc, ι_desc_assoc, + ι_colimMap_assoc, Discrete.functor_obj_eq_as, Discrete.natTrans_app, colimit.ι_desc_assoc, + Cofan.mk_pt, Cofan.mk_ι_app, ι_π, ι_π_assoc] + split + all_goals aesop + rw [this] + infer_instance + +instance Pi.map_epi {f g : J → C} [HasBiproduct f] [HasBiproduct g] (p : ∀ j, f j ⟶ g j) + [∀ j, Epi (p j)] : Epi (Pi.map p) := by + rw [show Pi.map p = (biproduct.isoProduct _).inv ≫ biproduct.map p ≫ + (biproduct.isoProduct _).hom by aesop] + infer_instance + +instance biproduct.map_mono {f g : J → C} [HasBiproduct f] [HasBiproduct g] (p : ∀ j, f j ⟶ g j) + [∀ j, Mono (p j)] : Mono (biproduct.map p) := by + rw [show biproduct.map p = (biproduct.isoProduct _).hom ≫ Pi.map p ≫ + (biproduct.isoProduct _).inv by aesop] + infer_instance + +instance Sigma.map_mono {f g : J → C} [HasBiproduct f] [HasBiproduct g] (p : ∀ j, f j ⟶ g j) + [∀ j, Mono (p j)] : Mono (Sigma.map p) := by + rw [show Sigma.map p = (biproduct.isoCoproduct _).inv ≫ biproduct.map p ≫ + (biproduct.isoCoproduct _).hom by aesop] + infer_instance + /-- Two biproducts which differ by an equivalence in the indexing type, and up to isomorphism in the factors, are isomorphic. @@ -1702,6 +1733,30 @@ theorem biprod.isIso_inl_iff_id_eq_fst_comp_inl (X Y : C) [HasBinaryBiproduct X · intro h exact ⟨⟨biprod.fst, biprod.inl_fst, h.symm⟩⟩ +instance biprod.map_epi {W X Y Z : C} (f : W ⟶ Y) (g : X ⟶ Z) [Epi f] + [Epi g] [HasBinaryBiproduct W X] [HasBinaryBiproduct Y Z] : Epi (biprod.map f g) := by + rw [show biprod.map f g = + (biprod.isoCoprod _ _).hom ≫ coprod.map f g ≫ (biprod.isoCoprod _ _).inv by aesop] + infer_instance + +instance prod.map_epi {W X Y Z : C} (f : W ⟶ Y) (g : X ⟶ Z) [Epi f] + [Epi g] [HasBinaryBiproduct W X] [HasBinaryBiproduct Y Z] : Epi (prod.map f g) := by + rw [show prod.map f g = (biprod.isoProd _ _).inv ≫ biprod.map f g ≫ + (biprod.isoProd _ _).hom by aesop] + infer_instance + +instance biprod.map_mono {W X Y Z : C} (f : W ⟶ Y) (g : X ⟶ Z) [Mono f] + [Mono g] [HasBinaryBiproduct W X] [HasBinaryBiproduct Y Z] : Mono (biprod.map f g) := by + rw [show biprod.map f g = (biprod.isoProd _ _).hom ≫ prod.map f g ≫ + (biprod.isoProd _ _).inv by aesop] + infer_instance + +instance coprod.map_mono {W X Y Z : C} (f : W ⟶ Y) (g : X ⟶ Z) [Mono f] + [Mono g] [HasBinaryBiproduct W X] [HasBinaryBiproduct Y Z] : Mono (coprod.map f g) := by + rw [show coprod.map f g = (biprod.isoCoprod _ _).inv ≫ biprod.map f g ≫ + (biprod.isoCoprod _ _).hom by aesop] + infer_instance + section BiprodKernel section BinaryBicone From a0408b3fc3743b87da02c4b6078c22177cc2e08b Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 18 Nov 2024 13:41:04 +0000 Subject: [PATCH 067/829] chore(MeasureTheory): move `Measure.comap` to a new file (#19178) --- Mathlib.lean | 1 + Mathlib/MeasureTheory/Measure/Comap.lean | 159 ++++++++++++++++++ .../MeasureTheory/Measure/MeasureSpace.lean | 124 +------------- Mathlib/MeasureTheory/Measure/Restrict.lean | 2 +- 4 files changed, 162 insertions(+), 124 deletions(-) create mode 100644 Mathlib/MeasureTheory/Measure/Comap.lean diff --git a/Mathlib.lean b/Mathlib.lean index 8aba1c993e563..3e01bb608398d 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3600,6 +3600,7 @@ import Mathlib.MeasureTheory.MeasurableSpace.Prod import Mathlib.MeasureTheory.Measure.AEDisjoint import Mathlib.MeasureTheory.Measure.AEMeasurable import Mathlib.MeasureTheory.Measure.AddContent +import Mathlib.MeasureTheory.Measure.Comap import Mathlib.MeasureTheory.Measure.Complex import Mathlib.MeasureTheory.Measure.Content import Mathlib.MeasureTheory.Measure.ContinuousPreimage diff --git a/Mathlib/MeasureTheory/Measure/Comap.lean b/Mathlib/MeasureTheory/Measure/Comap.lean new file mode 100644 index 0000000000000..adefd18c588f4 --- /dev/null +++ b/Mathlib/MeasureTheory/Measure/Comap.lean @@ -0,0 +1,159 @@ +/- +Copyright (c) 2020 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov, Rémy Degenne +-/ +import Mathlib.MeasureTheory.Measure.MeasureSpace + +/-! +# Pullback of a measure + +In this file we define the pullback `MeasureTheory.Measure.comap f μ` +of a measure `μ` along an injective map `f` +such that the image of any measurable set under `f` is a null-measurable set. +If `f` does not have these properties, then we define `comap f μ` to be zero. + +In the future, we may decide to redefine `comap f μ` so that it gives meaningful results, e.g., +for covering maps like `(↑) : ℝ → AddCircle (1 : ℝ)`. +-/ + +open Function Set Filter +open scoped ENNReal + +noncomputable section + +namespace MeasureTheory + +namespace Measure + +variable {α β : Type*} {s : Set α} + +open Classical in +/-- Pullback of a `Measure` as a linear map. If `f` sends each measurable set to a measurable +set, then for each measurable set `s` we have `comapₗ f μ s = μ (f '' s)`. + +Note that if `f` is not injective, this definition assigns `Set.univ` measure zero. + +If the linearity is not needed, please use `comap` instead, which works for a larger class of +functions. `comapₗ` is an auxiliary definition and most lemmas deal with comap. -/ +def comapₗ [MeasurableSpace α] [MeasurableSpace β] (f : α → β) : Measure β →ₗ[ℝ≥0∞] Measure α := + if hf : Injective f ∧ ∀ s, MeasurableSet s → MeasurableSet (f '' s) then + liftLinear (OuterMeasure.comap f) fun μ s hs t => by + simp only [OuterMeasure.comap_apply, image_inter hf.1, image_diff hf.1] + apply le_toOuterMeasure_caratheodory + exact hf.2 s hs + else 0 + +theorem comapₗ_apply {_ : MeasurableSpace α} {_ : MeasurableSpace β} (f : α → β) + (hfi : Injective f) (hf : ∀ s, MeasurableSet s → MeasurableSet (f '' s)) (μ : Measure β) + (hs : MeasurableSet s) : comapₗ f μ s = μ (f '' s) := by + rw [comapₗ, dif_pos, liftLinear_apply _ hs, OuterMeasure.comap_apply, coe_toOuterMeasure] + exact ⟨hfi, hf⟩ + +open Classical in +/-- Pullback of a `Measure`. If `f` sends each measurable set to a null-measurable set, +then for each measurable set `s` we have `comap f μ s = μ (f '' s)`. + +Note that if `f` is not injective, this definition assigns `Set.univ` measure zero. -/ +def comap [MeasurableSpace α] [MeasurableSpace β] (f : α → β) (μ : Measure β) : Measure α := + if hf : Injective f ∧ ∀ s, MeasurableSet s → NullMeasurableSet (f '' s) μ then + (OuterMeasure.comap f μ.toOuterMeasure).toMeasure fun s hs t => by + simp only [OuterMeasure.comap_apply, image_inter hf.1, image_diff hf.1] + exact (measure_inter_add_diff₀ _ (hf.2 s hs)).symm + else 0 + +variable {ma : MeasurableSpace α} {mb : MeasurableSpace β} + +theorem comap_apply₀ (f : α → β) (μ : Measure β) (hfi : Injective f) + (hf : ∀ s, MeasurableSet s → NullMeasurableSet (f '' s) μ) + (hs : NullMeasurableSet s (comap f μ)) : comap f μ s = μ (f '' s) := by + rw [comap, dif_pos (And.intro hfi hf)] at hs ⊢ + rw [toMeasure_apply₀ _ _ hs, OuterMeasure.comap_apply, coe_toOuterMeasure] + +theorem le_comap_apply (f : α → β) (μ : Measure β) (hfi : Injective f) + (hf : ∀ s, MeasurableSet s → NullMeasurableSet (f '' s) μ) (s : Set α) : + μ (f '' s) ≤ comap f μ s := by + rw [comap, dif_pos (And.intro hfi hf)] + exact le_toMeasure_apply _ _ _ + +theorem comap_apply (f : α → β) (hfi : Injective f) + (hf : ∀ s, MeasurableSet s → MeasurableSet (f '' s)) (μ : Measure β) (hs : MeasurableSet s) : + comap f μ s = μ (f '' s) := + comap_apply₀ f μ hfi (fun s hs => (hf s hs).nullMeasurableSet) hs.nullMeasurableSet + +theorem comapₗ_eq_comap (f : α → β) (hfi : Injective f) + (hf : ∀ s, MeasurableSet s → MeasurableSet (f '' s)) (μ : Measure β) (hs : MeasurableSet s) : + comapₗ f μ s = comap f μ s := + (comapₗ_apply f hfi hf μ hs).trans (comap_apply f hfi hf μ hs).symm + +theorem measure_image_eq_zero_of_comap_eq_zero (f : α → β) (μ : Measure β) (hfi : Injective f) + (hf : ∀ s, MeasurableSet s → NullMeasurableSet (f '' s) μ) {s : Set α} (hs : comap f μ s = 0) : + μ (f '' s) = 0 := + le_antisymm ((le_comap_apply f μ hfi hf s).trans hs.le) (zero_le _) + +theorem ae_eq_image_of_ae_eq_comap (f : α → β) (μ : Measure β) (hfi : Injective f) + (hf : ∀ s, MeasurableSet s → NullMeasurableSet (f '' s) μ) + {s t : Set α} (hst : s =ᵐ[comap f μ] t) : f '' s =ᵐ[μ] f '' t := by + rw [EventuallyEq, ae_iff] at hst ⊢ + have h_eq_α : { a : α | ¬s a = t a } = s \ t ∪ t \ s := by + ext1 x + simp only [eq_iff_iff, mem_setOf_eq, mem_union, mem_diff] + tauto + have h_eq_β : { a : β | ¬(f '' s) a = (f '' t) a } = f '' s \ f '' t ∪ f '' t \ f '' s := by + ext1 x + simp only [eq_iff_iff, mem_setOf_eq, mem_union, mem_diff] + tauto + rw [← Set.image_diff hfi, ← Set.image_diff hfi, ← Set.image_union] at h_eq_β + rw [h_eq_β] + rw [h_eq_α] at hst + exact measure_image_eq_zero_of_comap_eq_zero f μ hfi hf hst + +theorem NullMeasurableSet.image (f : α → β) (μ : Measure β) (hfi : Injective f) + (hf : ∀ s, MeasurableSet s → NullMeasurableSet (f '' s) μ) + (hs : NullMeasurableSet s (μ.comap f)) : NullMeasurableSet (f '' s) μ := by + refine ⟨toMeasurable μ (f '' toMeasurable (μ.comap f) s), measurableSet_toMeasurable _ _, ?_⟩ + refine EventuallyEq.trans ?_ (NullMeasurableSet.toMeasurable_ae_eq ?_).symm + swap + · exact hf _ (measurableSet_toMeasurable _ _) + have h : toMeasurable (comap f μ) s =ᵐ[comap f μ] s := + NullMeasurableSet.toMeasurable_ae_eq hs + exact ae_eq_image_of_ae_eq_comap f μ hfi hf h.symm + +theorem comap_preimage (f : α → β) (μ : Measure β) (hf : Injective f) (hf' : Measurable f) + (h : ∀ t, MeasurableSet t → NullMeasurableSet (f '' t) μ) {s : Set β} (hs : MeasurableSet s) : + μ.comap f (f ⁻¹' s) = μ (s ∩ range f) := by + rw [comap_apply₀ _ _ hf h (hf' hs).nullMeasurableSet, image_preimage_eq_inter_range] + +@[simp] lemma comap_zero (f : α → β) : (0 : Measure β).comap f = 0 := by + by_cases hf : Injective f ∧ ∀ s, MeasurableSet s → NullMeasurableSet (f '' s) (0 : Measure β) + · simp [comap, hf] + · simp [comap, hf] + +end Measure + +end MeasureTheory + +open MeasureTheory Measure + +variable {α β : Type*} {ma : MeasurableSpace α} {mb : MeasurableSpace β} + +lemma MeasurableEmbedding.comap_add {f : α → β} (hf : MeasurableEmbedding f) (μ ν : Measure β) : + (μ + ν).comap f = μ.comap f + ν.comap f := by + ext s hs + simp only [← comapₗ_eq_comap _ hf.injective (fun _ ↦ hf.measurableSet_image.mpr) _ hs, + _root_.map_add, add_apply] + +namespace MeasurableEquiv + +lemma comap_symm {μ : Measure α} (e : α ≃ᵐ β) : μ.comap e.symm = μ.map e := by + ext s hs + rw [e.map_apply, Measure.comap_apply _ e.symm.injective _ _ hs, image_symm] + exact fun t ht ↦ e.symm.measurableSet_image.mpr ht + +lemma map_symm {μ : Measure α} (e : β ≃ᵐ α) : μ.map e.symm = μ.comap e := by + rw [← comap_symm, symm_symm] + +end MeasurableEquiv + +lemma comap_swap (μ : Measure (α × β)) : μ.comap Prod.swap = μ.map Prod.swap := + (MeasurableEquiv.prodComm ..).comap_symm diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index e80720475f839..754db225acc6d 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -1244,107 +1244,6 @@ theorem preimage_null_of_map_null {f : α → β} (hf : AEMeasurable f μ) {s : theorem tendsto_ae_map {f : α → β} (hf : AEMeasurable f μ) : Tendsto f (ae μ) (ae (μ.map f)) := fun _ hs => preimage_null_of_map_null hf hs -open Classical in -/-- Pullback of a `Measure` as a linear map. If `f` sends each measurable set to a measurable -set, then for each measurable set `s` we have `comapₗ f μ s = μ (f '' s)`. - -Note that if `f` is not injective, this definition assigns `Set.univ` measure zero. - -If the linearity is not needed, please use `comap` instead, which works for a larger class of -functions. `comapₗ` is an auxiliary definition and most lemmas deal with comap. -/ -def comapₗ [MeasurableSpace α] (f : α → β) : Measure β →ₗ[ℝ≥0∞] Measure α := - if hf : Injective f ∧ ∀ s, MeasurableSet s → MeasurableSet (f '' s) then - liftLinear (OuterMeasure.comap f) fun μ s hs t => by - simp only [OuterMeasure.comap_apply, image_inter hf.1, image_diff hf.1] - apply le_toOuterMeasure_caratheodory - exact hf.2 s hs - else 0 - -theorem comapₗ_apply {β} {_ : MeasurableSpace α} {mβ : MeasurableSpace β} (f : α → β) - (hfi : Injective f) (hf : ∀ s, MeasurableSet s → MeasurableSet (f '' s)) (μ : Measure β) - (hs : MeasurableSet s) : comapₗ f μ s = μ (f '' s) := by - rw [comapₗ, dif_pos, liftLinear_apply _ hs, OuterMeasure.comap_apply, coe_toOuterMeasure] - exact ⟨hfi, hf⟩ - -open Classical in -/-- Pullback of a `Measure`. If `f` sends each measurable set to a null-measurable set, -then for each measurable set `s` we have `comap f μ s = μ (f '' s)`. - -Note that if `f` is not injective, this definition assigns `Set.univ` measure zero. -/ -def comap [MeasurableSpace α] (f : α → β) (μ : Measure β) : Measure α := - if hf : Injective f ∧ ∀ s, MeasurableSet s → NullMeasurableSet (f '' s) μ then - (OuterMeasure.comap f μ.toOuterMeasure).toMeasure fun s hs t => by - simp only [OuterMeasure.comap_apply, image_inter hf.1, image_diff hf.1] - exact (measure_inter_add_diff₀ _ (hf.2 s hs)).symm - else 0 - -theorem comap_apply₀ {_ : MeasurableSpace α} (f : α → β) (μ : Measure β) (hfi : Injective f) - (hf : ∀ s, MeasurableSet s → NullMeasurableSet (f '' s) μ) - (hs : NullMeasurableSet s (comap f μ)) : comap f μ s = μ (f '' s) := by - rw [comap, dif_pos (And.intro hfi hf)] at hs ⊢ - rw [toMeasure_apply₀ _ _ hs, OuterMeasure.comap_apply, coe_toOuterMeasure] - -theorem le_comap_apply {β} {_ : MeasurableSpace α} {mβ : MeasurableSpace β} (f : α → β) - (μ : Measure β) (hfi : Injective f) (hf : ∀ s, MeasurableSet s → NullMeasurableSet (f '' s) μ) - (s : Set α) : μ (f '' s) ≤ comap f μ s := by - rw [comap, dif_pos (And.intro hfi hf)] - exact le_toMeasure_apply _ _ _ - -theorem comap_apply {β} {_ : MeasurableSpace α} {_mβ : MeasurableSpace β} (f : α → β) - (hfi : Injective f) (hf : ∀ s, MeasurableSet s → MeasurableSet (f '' s)) (μ : Measure β) - (hs : MeasurableSet s) : comap f μ s = μ (f '' s) := - comap_apply₀ f μ hfi (fun s hs => (hf s hs).nullMeasurableSet) hs.nullMeasurableSet - -theorem comapₗ_eq_comap {β} {_ : MeasurableSpace α} {_mβ : MeasurableSpace β} (f : α → β) - (hfi : Injective f) (hf : ∀ s, MeasurableSet s → MeasurableSet (f '' s)) (μ : Measure β) - (hs : MeasurableSet s) : comapₗ f μ s = comap f μ s := - (comapₗ_apply f hfi hf μ hs).trans (comap_apply f hfi hf μ hs).symm - -theorem measure_image_eq_zero_of_comap_eq_zero {β} {_ : MeasurableSpace α} {_mβ : MeasurableSpace β} - (f : α → β) (μ : Measure β) (hfi : Injective f) - (hf : ∀ s, MeasurableSet s → NullMeasurableSet (f '' s) μ) {s : Set α} (hs : comap f μ s = 0) : - μ (f '' s) = 0 := - le_antisymm ((le_comap_apply f μ hfi hf s).trans hs.le) (zero_le _) - -theorem ae_eq_image_of_ae_eq_comap {β} {_ : MeasurableSpace α} {mβ : MeasurableSpace β} (f : α → β) - (μ : Measure β) (hfi : Injective f) (hf : ∀ s, MeasurableSet s → NullMeasurableSet (f '' s) μ) - {s t : Set α} (hst : s =ᵐ[comap f μ] t) : f '' s =ᵐ[μ] f '' t := by - rw [EventuallyEq, ae_iff] at hst ⊢ - have h_eq_α : { a : α | ¬s a = t a } = s \ t ∪ t \ s := by - ext1 x - simp only [eq_iff_iff, mem_setOf_eq, mem_union, mem_diff] - tauto - have h_eq_β : { a : β | ¬(f '' s) a = (f '' t) a } = f '' s \ f '' t ∪ f '' t \ f '' s := by - ext1 x - simp only [eq_iff_iff, mem_setOf_eq, mem_union, mem_diff] - tauto - rw [← Set.image_diff hfi, ← Set.image_diff hfi, ← Set.image_union] at h_eq_β - rw [h_eq_β] - rw [h_eq_α] at hst - exact measure_image_eq_zero_of_comap_eq_zero f μ hfi hf hst - -theorem NullMeasurableSet.image {β} {_ : MeasurableSpace α} {mβ : MeasurableSpace β} (f : α → β) - (μ : Measure β) (hfi : Injective f) (hf : ∀ s, MeasurableSet s → NullMeasurableSet (f '' s) μ) - {s : Set α} (hs : NullMeasurableSet s (μ.comap f)) : NullMeasurableSet (f '' s) μ := by - refine ⟨toMeasurable μ (f '' toMeasurable (μ.comap f) s), measurableSet_toMeasurable _ _, ?_⟩ - refine EventuallyEq.trans ?_ (NullMeasurableSet.toMeasurable_ae_eq ?_).symm - swap - · exact hf _ (measurableSet_toMeasurable _ _) - have h : toMeasurable (comap f μ) s =ᵐ[comap f μ] s := - NullMeasurableSet.toMeasurable_ae_eq hs - exact ae_eq_image_of_ae_eq_comap f μ hfi hf h.symm - -theorem comap_preimage {β} {_ : MeasurableSpace α} {mβ : MeasurableSpace β} (f : α → β) - (μ : Measure β) {s : Set β} (hf : Injective f) (hf' : Measurable f) - (h : ∀ t, MeasurableSet t → NullMeasurableSet (f '' t) μ) (hs : MeasurableSet s) : - μ.comap f (f ⁻¹' s) = μ (s ∩ range f) := by - rw [comap_apply₀ _ _ hf h (hf' hs).nullMeasurableSet, image_preimage_eq_inter_range] - -@[simp] lemma comap_zero : (0 : Measure β).comap f = 0 := by - by_cases hf : Injective f ∧ ∀ s, MeasurableSet s → NullMeasurableSet (f '' s) (0 : Measure β) - · simp [comap, hf] - · simp [comap, hf] - section Sum variable {f : ι → Measure α} @@ -1953,12 +1852,6 @@ nonrec theorem map_apply (hf : MeasurableEmbedding f) (μ : Measure α) (s : Set μ.map f s ≤ μ.map f t := by gcongr _ = μ (f ⁻¹' s) := by rw [map_apply hf.measurable htm, hft, measure_toMeasurable] -lemma comap_add (hf : MeasurableEmbedding f) (μ ν : Measure β) : - (μ + ν).comap f = μ.comap f + ν.comap f := by - ext s hs - simp only [← comapₗ_eq_comap _ hf.injective (fun _ ↦ hf.measurableSet_image.mpr) _ hs, - _root_.map_add, add_apply] - lemma absolutelyContinuous_map (hf : MeasurableEmbedding f) (hμν : μ ≪ ν) : μ.map f ≪ ν.map f := by intro t ht @@ -1980,14 +1873,6 @@ variable {_ : MeasurableSpace α} [MeasurableSpace β] {μ : Measure α} {ν : M protected theorem map_apply (f : α ≃ᵐ β) (s : Set β) : μ.map f s = μ (f ⁻¹' s) := f.measurableEmbedding.map_apply _ _ -lemma comap_symm (e : α ≃ᵐ β) : μ.comap e.symm = μ.map e := by - ext s hs - rw [e.map_apply, Measure.comap_apply _ e.symm.injective _ _ hs, image_symm] - exact fun t ht ↦ e.symm.measurableSet_image.mpr ht - -lemma map_symm (e : β ≃ᵐ α) : μ.map e.symm = μ.comap e := by - rw [← comap_symm, symm_symm] - @[simp] theorem map_symm_map (e : α ≃ᵐ β) : (μ.map e).map e.symm = μ := by simp [map_map e.symm.measurable e.measurable] @@ -2014,13 +1899,6 @@ theorem quasiMeasurePreserving_symm (μ : Measure α) (e : α ≃ᵐ β) : end MeasurableEquiv -namespace MeasureTheory.Measure -variable {mα : MeasurableSpace α} {mβ : MeasurableSpace β} - -lemma comap_swap (μ : Measure (α × β)) : μ.comap Prod.swap = μ.map Prod.swap := - (MeasurableEquiv.prodComm ..).comap_symm - -end MeasureTheory.Measure end -set_option linter.style.longFile 2200 +set_option linter.style.longFile 2100 diff --git a/Mathlib/MeasureTheory/Measure/Restrict.lean b/Mathlib/MeasureTheory/Measure/Restrict.lean index f9ddf8fd820c1..b3823303c8980 100644 --- a/Mathlib/MeasureTheory/Measure/Restrict.lean +++ b/Mathlib/MeasureTheory/Measure/Restrict.lean @@ -3,7 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ -import Mathlib.MeasureTheory.Measure.MeasureSpace +import Mathlib.MeasureTheory.Measure.Comap /-! # Restricting a measure to a subset or a subtype From 2f99cb3a8fd8419e39c596237219e8867be6deea Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 18 Nov 2024 13:57:16 +0000 Subject: [PATCH 068/829] chore: make sigma-algebra implicit in the `Measure.map` API (#19026) This sigma-algebra can be inferred from context and will be filled in non-canonically in the context of Gibbs measures. From GibbsMeasure --- .../MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean | 4 ++-- Mathlib/MeasureTheory/Measure/MeasureSpace.lean | 9 +++++---- Mathlib/Probability/Kernel/Condexp.lean | 4 ++-- 3 files changed, 9 insertions(+), 8 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean index 8cde3b7093566..be988e8efe5a8 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean @@ -96,7 +96,7 @@ theorem MeasureTheory.measure_lt_one_eq_integral_div_gamma {p : ℝ} (hp : 0 < p (LinearEquiv.refl ℝ E : E ≃ₗ[ℝ] F) -- The measure `ν` is the measure on `F` defined by `μ` -- Since we have two different topologies, it is necessary to specify the topology of E - let ν : Measure F := @Measure.map E F _ mE φ μ + let ν : Measure F := @Measure.map E F mE _ φ μ have : IsAddHaarMeasure ν := @ContinuousLinearEquiv.isAddHaarMeasure_map E F ℝ ℝ _ _ _ _ _ _ tE _ _ _ _ _ _ _ mE _ _ _ φ μ _ convert (measure_unitBall_eq_integral_div_gamma ν hp) using 1 @@ -139,7 +139,7 @@ theorem MeasureTheory.measure_le_eq_lt [Nontrivial E] (r : ℝ) : (LinearEquiv.refl ℝ E : E ≃ₗ[ℝ] F) -- The measure `ν` is the measure on `F` defined by `μ` -- Since we have two different topologies, it is necessary to specify the topology of E - let ν : Measure F := @Measure.map E F _ mE φ μ + let ν : Measure F := @Measure.map E F mE _ φ μ have : IsAddHaarMeasure ν := @ContinuousLinearEquiv.isAddHaarMeasure_map E F ℝ ℝ _ _ _ _ _ _ tE _ _ _ _ _ _ _ mE _ _ _ φ μ _ convert addHaar_closedBall_eq_addHaar_ball ν 0 r using 1 diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index 754db225acc6d..8cf14fc5dc9a6 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -698,7 +698,7 @@ section /- Porting note: These variables are wrapped by an anonymous section because they interrupt synthesizing instances in `MeasureSpace` section. -/ -variable {m0 : MeasurableSpace α} [MeasurableSpace β] [MeasurableSpace γ] +variable {m0 : MeasurableSpace α} {mβ : MeasurableSpace β} [MeasurableSpace γ] variable {μ μ₁ μ₂ μ₃ ν ν' ν₁ ν₂ : Measure α} {s s' t : Set α} namespace Measure @@ -1058,7 +1058,7 @@ lemma nonempty_of_neZero (μ : Measure α) [NeZero μ] : Nonempty α := /-- Lift a linear map between `OuterMeasure` spaces such that for each measure `μ` every measurable set is caratheodory-measurable w.r.t. `f μ` to a linear map between `Measure` spaces. -/ -def liftLinear {m0 : MeasurableSpace α} (f : OuterMeasure α →ₗ[ℝ≥0∞] OuterMeasure β) +def liftLinear [MeasurableSpace β] (f : OuterMeasure α →ₗ[ℝ≥0∞] OuterMeasure β) (hf : ∀ μ : Measure α, ‹_› ≤ (f μ.toOuterMeasure).caratheodory) : Measure α →ₗ[ℝ≥0∞] Measure β where toFun μ := (f μ.toOuterMeasure).toMeasure (hf μ) @@ -1086,7 +1086,7 @@ theorem le_liftLinear_apply {f : OuterMeasure α →ₗ[ℝ≥0∞] OuterMeasure open Classical in /-- The pushforward of a measure as a linear map. It is defined to be `0` if `f` is not a measurable function. -/ -def mapₗ [MeasurableSpace α] (f : α → β) : Measure α →ₗ[ℝ≥0∞] Measure β := +def mapₗ [MeasurableSpace α] [MeasurableSpace β] (f : α → β) : Measure α →ₗ[ℝ≥0∞] Measure β := if hf : Measurable f then liftLinear (OuterMeasure.map f) fun μ _s hs t => le_toOuterMeasure_caratheodory μ _ (hf hs) (f ⁻¹' t) @@ -1101,7 +1101,8 @@ theorem mapₗ_congr {f g : α → β} (hf : Measurable f) (hg : Measurable g) ( open Classical in /-- The pushforward of a measure. It is defined to be `0` if `f` is not an almost everywhere measurable function. -/ -irreducible_def map [MeasurableSpace α] (f : α → β) (μ : Measure α) : Measure β := +irreducible_def map [MeasurableSpace α] [MeasurableSpace β] (f : α → β) (μ : Measure α) : + Measure β := if hf : AEMeasurable f μ then mapₗ (hf.mk f) μ else 0 theorem mapₗ_mk_apply_of_aemeasurable {f : α → β} (hf : AEMeasurable f μ) : diff --git a/Mathlib/Probability/Kernel/Condexp.lean b/Mathlib/Probability/Kernel/Condexp.lean index 6981a66d56ad7..f2b085aafd5f5 100644 --- a/Mathlib/Probability/Kernel/Condexp.lean +++ b/Mathlib/Probability/Kernel/Condexp.lean @@ -42,14 +42,14 @@ variable {Ω F : Type*} {m mΩ : MeasurableSpace Ω} {μ : Measure Ω} {f : Ω theorem _root_.MeasureTheory.AEStronglyMeasurable.comp_snd_map_prod_id [TopologicalSpace F] (hm : m ≤ mΩ) (hf : AEStronglyMeasurable f μ) : AEStronglyMeasurable (fun x : Ω × Ω => f x.2) - (@Measure.map Ω (Ω × Ω) (m.prod mΩ) mΩ (fun ω => (id ω, id ω)) μ) := by + (@Measure.map Ω (Ω × Ω) mΩ (m.prod mΩ) (fun ω => (id ω, id ω)) μ) := by rw [← aestronglyMeasurable_comp_snd_map_prod_mk_iff (measurable_id'' hm)] at hf simp_rw [id] at hf ⊢ exact hf theorem _root_.MeasureTheory.Integrable.comp_snd_map_prod_id [NormedAddCommGroup F] (hm : m ≤ mΩ) (hf : Integrable f μ) : Integrable (fun x : Ω × Ω => f x.2) - (@Measure.map Ω (Ω × Ω) (m.prod mΩ) mΩ (fun ω => (id ω, id ω)) μ) := by + (@Measure.map Ω (Ω × Ω) mΩ (m.prod mΩ) (fun ω => (id ω, id ω)) μ) := by rw [← integrable_comp_snd_map_prod_mk_iff (measurable_id'' hm)] at hf simp_rw [id] at hf ⊢ exact hf From aed31ba4345ad7023e0a4ee8ae3ed70dbac9b773 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 18 Nov 2024 13:57:18 +0000 Subject: [PATCH 069/829] chore(MeasureTheory): move 2 files about vector-valued measures to a new folder (#19168) --- Mathlib.lean | 4 ++-- Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean | 2 +- Mathlib/MeasureTheory/Decomposition/SignedHahn.lean | 2 +- Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean | 2 +- Mathlib/MeasureTheory/Measure/Complex.lean | 2 +- .../{Measure/VectorMeasure.lean => VectorMeasure/Basic.lean} | 0 .../WithDensity.lean} | 2 +- 7 files changed, 7 insertions(+), 7 deletions(-) rename Mathlib/MeasureTheory/{Measure/VectorMeasure.lean => VectorMeasure/Basic.lean} (100%) rename Mathlib/MeasureTheory/{Measure/WithDensityVectorMeasure.lean => VectorMeasure/WithDensity.lean} (99%) diff --git a/Mathlib.lean b/Mathlib.lean index 3e01bb608398d..c6aea2fa69b6f 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3644,10 +3644,8 @@ import Mathlib.MeasureTheory.Measure.Sub import Mathlib.MeasureTheory.Measure.Tilted import Mathlib.MeasureTheory.Measure.Trim import Mathlib.MeasureTheory.Measure.Typeclasses -import Mathlib.MeasureTheory.Measure.VectorMeasure import Mathlib.MeasureTheory.Measure.WithDensity import Mathlib.MeasureTheory.Measure.WithDensityFinite -import Mathlib.MeasureTheory.Measure.WithDensityVectorMeasure import Mathlib.MeasureTheory.Order.Group.Lattice import Mathlib.MeasureTheory.Order.Lattice import Mathlib.MeasureTheory.Order.UpperLower @@ -3662,6 +3660,8 @@ import Mathlib.MeasureTheory.OuterMeasure.Operations import Mathlib.MeasureTheory.PiSystem import Mathlib.MeasureTheory.SetAlgebra import Mathlib.MeasureTheory.SetSemiring +import Mathlib.MeasureTheory.VectorMeasure.Basic +import Mathlib.MeasureTheory.VectorMeasure.WithDensity import Mathlib.ModelTheory.Algebra.Field.Basic import Mathlib.ModelTheory.Algebra.Field.CharP import Mathlib.ModelTheory.Algebra.Field.IsAlgClosed diff --git a/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean b/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean index 088ab54bf263c..d0c1bf8f1d942 100644 --- a/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean +++ b/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kexing Ying, Rémy Degenne -/ import Mathlib.MeasureTheory.Decomposition.SignedLebesgue -import Mathlib.MeasureTheory.Measure.WithDensityVectorMeasure +import Mathlib.MeasureTheory.VectorMeasure.WithDensity /-! # Radon-Nikodym theorem diff --git a/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean b/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean index 77cd7b8a79759..f6fe47d6c9795 100644 --- a/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean +++ b/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Kexing Ying. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kexing Ying -/ -import Mathlib.MeasureTheory.Measure.VectorMeasure +import Mathlib.MeasureTheory.VectorMeasure.Basic import Mathlib.Order.SymmDiff /-! diff --git a/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean b/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean index 7652993309eaf..354c4e16c0210 100644 --- a/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean +++ b/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean @@ -6,7 +6,7 @@ Authors: Kexing Ying import Mathlib.MeasureTheory.Decomposition.Lebesgue import Mathlib.MeasureTheory.Measure.Complex import Mathlib.MeasureTheory.Decomposition.Jordan -import Mathlib.MeasureTheory.Measure.WithDensityVectorMeasure +import Mathlib.MeasureTheory.VectorMeasure.WithDensity /-! # Lebesgue decomposition diff --git a/Mathlib/MeasureTheory/Measure/Complex.lean b/Mathlib/MeasureTheory/Measure/Complex.lean index 611174a941c5f..a2f8f37c1c3a4 100644 --- a/Mathlib/MeasureTheory/Measure/Complex.lean +++ b/Mathlib/MeasureTheory/Measure/Complex.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Kexing Ying. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kexing Ying -/ -import Mathlib.MeasureTheory.Measure.VectorMeasure +import Mathlib.MeasureTheory.VectorMeasure.Basic import Mathlib.Analysis.Complex.Basic /-! diff --git a/Mathlib/MeasureTheory/Measure/VectorMeasure.lean b/Mathlib/MeasureTheory/VectorMeasure/Basic.lean similarity index 100% rename from Mathlib/MeasureTheory/Measure/VectorMeasure.lean rename to Mathlib/MeasureTheory/VectorMeasure/Basic.lean diff --git a/Mathlib/MeasureTheory/Measure/WithDensityVectorMeasure.lean b/Mathlib/MeasureTheory/VectorMeasure/WithDensity.lean similarity index 99% rename from Mathlib/MeasureTheory/Measure/WithDensityVectorMeasure.lean rename to Mathlib/MeasureTheory/VectorMeasure/WithDensity.lean index 095d95dd608a0..e0642d458ddf3 100644 --- a/Mathlib/MeasureTheory/Measure/WithDensityVectorMeasure.lean +++ b/Mathlib/MeasureTheory/VectorMeasure/WithDensity.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Kexing Ying. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kexing Ying -/ -import Mathlib.MeasureTheory.Measure.VectorMeasure +import Mathlib.MeasureTheory.VectorMeasure.Basic import Mathlib.MeasureTheory.Function.AEEqOfIntegral /-! From 941de860cd992810ce3e7a9281ce1340ff098e04 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Mon, 18 Nov 2024 14:41:27 +0000 Subject: [PATCH 070/829] chore: all variants of `limitOpIsoOpColimit` (#18558) --- Mathlib/CategoryTheory/Limits/IndYoneda.lean | 39 ------ Mathlib/CategoryTheory/Limits/Opposites.lean | 120 +++++++++++++++++++ 2 files changed, 120 insertions(+), 39 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/IndYoneda.lean b/Mathlib/CategoryTheory/Limits/IndYoneda.lean index 85bb65a0003ad..45d56e6704af1 100644 --- a/Mathlib/CategoryTheory/Limits/IndYoneda.lean +++ b/Mathlib/CategoryTheory/Limits/IndYoneda.lean @@ -34,45 +34,6 @@ open Opposite variable {C : Type u₁} [Category.{u₂} C] {I : Type v₁} [Category.{v₂} I] -section LimitOpColimit - -/-- The limit of `F.op` is the opposite of `colimit F`. -/ -noncomputable def limitOpIsoOpColimit (F : I ⥤ C) [HasColimit F] : - limit F.op ≅ op <| colimit F := - limit.isoLimitCone ⟨(colimit.cocone F).op, (colimit.isColimit F).op⟩ - -@[reassoc (attr := simp)] -lemma limitOpIsoOpColimit_inv_comp_π (F : I ⥤ C) [HasColimit F] (i : I) : - (limitOpIsoOpColimit F).inv ≫ limit.π F.op ⟨i⟩ = (colimit.ι F i).op := by - simp [limitOpIsoOpColimit] - -@[reassoc (attr := simp)] -lemma limitOpIsoOpColimit_hom_comp_ι (F : I ⥤ C) [HasColimit F] (i : I) : - (limitOpIsoOpColimit F).hom ≫ (colimit.ι F i).op = limit.π F.op ⟨i⟩ := by - rw [← limitOpIsoOpColimit_inv_comp_π, ← Category.assoc, Iso.hom_inv_id, Category.id_comp] - -instance (F : Iᵒᵖ ⥤ C) [HasColimit F] : HasLimit F.rightOp where - exists_limit := ⟨⟨coneRightOpOfCocone (colimit.cocone F), - isLimitConeRightOpOfCocone F (colimit.isColimit F)⟩⟩ - -/-- `limitOpIsoOpColimit` for contravariant functor. -/ -noncomputable def limitRightOpIsoOpColimit (F : Iᵒᵖ ⥤ C) [HasColimit F] : - limit F.rightOp ≅ op <| colimit F := - limit.isoLimitCone ⟨coneRightOpOfCocone (colimit.cocone F), - isLimitConeRightOpOfCocone F (colimit.isColimit F)⟩ - -@[reassoc (attr := simp)] -lemma limitRightOpIsoOpColimit_inv_comp_π (F : Iᵒᵖ ⥤ C) [HasColimit F] (i : I) : - (limitRightOpIsoOpColimit F).inv ≫ limit.π F.rightOp i = (colimit.ι F ⟨i⟩).op := by - simp [limitRightOpIsoOpColimit] - -@[reassoc (attr := simp)] -lemma limitRightOpIsoOpColimit_hom_comp_ι (F : Iᵒᵖ ⥤ C) [HasColimit F] (i : I) : - (limitRightOpIsoOpColimit F).hom ≫ (colimit.ι F ⟨i⟩).op = limit.π F.rightOp i := by - rw [← limitRightOpIsoOpColimit_inv_comp_π, ← Category.assoc, Iso.hom_inv_id, Category.id_comp] - -end LimitOpColimit - section HomCocontinuousCovariant variable (F : I ⥤ C) [HasColimit F] diff --git a/Mathlib/CategoryTheory/Limits/Opposites.lean b/Mathlib/CategoryTheory/Limits/Opposites.lean index 58396d9c13e29..493dec55ff8dc 100644 --- a/Mathlib/CategoryTheory/Limits/Opposites.lean +++ b/Mathlib/CategoryTheory/Limits/Opposites.lean @@ -285,6 +285,66 @@ instance hasLimit_unop_of_hasColimit (F : Jᵒᵖ ⥤ Cᵒᵖ) [HasColimit F] : { cone := coneUnopOfCocone (colimit.cocone F) isLimit := isLimitConeUnopOfCocone _ (colimit.isColimit _) } +/-- The limit of `F.op` is the opposite of `colimit F`. -/ +def limitOpIsoOpColimit (F : J ⥤ C) [HasColimit F] : + limit F.op ≅ op (colimit F) := + limit.isoLimitCone ⟨_, (colimit.isColimit _).op⟩ + +@[reassoc (attr := simp)] +lemma limitOpIsoOpColimit_inv_comp_π (F : J ⥤ C) [HasColimit F] (j : Jᵒᵖ) : + (limitOpIsoOpColimit F).inv ≫ limit.π F.op j = (colimit.ι F j.unop).op := by + simp [limitOpIsoOpColimit] + +@[reassoc (attr := simp)] +lemma limitOpIsoOpColimit_hom_comp_ι (F : J ⥤ C) [HasColimit F] (j : J) : + (limitOpIsoOpColimit F).hom ≫ (colimit.ι F j).op = limit.π F.op (op j) := by + simp [← Iso.eq_inv_comp] + +/-- The limit of `F.leftOp` is the unopposite of `colimit F`. -/ +def limitLeftOpIsoUnopColimit (F : J ⥤ Cᵒᵖ) [HasColimit F] : + limit F.leftOp ≅ unop (colimit F) := + limit.isoLimitCone ⟨_, isLimitConeLeftOpOfCocone _ (colimit.isColimit _)⟩ + +@[reassoc (attr := simp)] +lemma limitLeftOpIsoUnopColimit_inv_comp_π (F : J ⥤ Cᵒᵖ) [HasColimit F] (j : Jᵒᵖ) : + (limitLeftOpIsoUnopColimit F).inv ≫ limit.π F.leftOp j = (colimit.ι F j.unop).unop := by + simp [limitLeftOpIsoUnopColimit] + +@[reassoc (attr := simp)] +lemma limitLeftOpIsoUnopColimit_hom_comp_ι (F : J ⥤ Cᵒᵖ) [HasColimit F] (j : J) : + (limitLeftOpIsoUnopColimit F).hom ≫ (colimit.ι F j).unop = limit.π F.leftOp (op j) := by + simp [← Iso.eq_inv_comp] + +/-- The limit of `F.rightOp` is the opposite of `colimit F`. -/ +def limitRightOpIsoOpColimit (F : Jᵒᵖ ⥤ C) [HasColimit F] : + limit F.rightOp ≅ op (colimit F) := + limit.isoLimitCone ⟨_, isLimitConeRightOpOfCocone _ (colimit.isColimit _)⟩ + +@[reassoc (attr := simp)] +lemma limitRightOpIsoOpColimit_inv_comp_π (F : Jᵒᵖ ⥤ C) [HasColimit F] (j : J) : + (limitRightOpIsoOpColimit F).inv ≫ limit.π F.rightOp j = (colimit.ι F (op j)).op := by + simp [limitRightOpIsoOpColimit] + +@[reassoc (attr := simp)] +lemma limitRightOpIsoOpColimit_hom_comp_ι (F : Jᵒᵖ ⥤ C) [HasColimit F] (j : Jᵒᵖ) : + (limitRightOpIsoOpColimit F).hom ≫ (colimit.ι F j).op = limit.π F.rightOp j.unop := by + simp [← Iso.eq_inv_comp] + +/-- The limit of `F.unop` is the unopposite of `colimit F`. -/ +def limitUnopIsoUnopColimit (F : Jᵒᵖ ⥤ Cᵒᵖ) [HasColimit F] : + limit F.unop ≅ unop (colimit F) := + limit.isoLimitCone ⟨_, isLimitConeUnopOfCocone _ (colimit.isColimit _)⟩ + +@[reassoc (attr := simp)] +lemma limitUnopIsoUnopColimit_inv_comp_π (F : Jᵒᵖ ⥤ Cᵒᵖ) [HasColimit F] (j : J) : + (limitUnopIsoUnopColimit F).inv ≫ limit.π F.unop j = (colimit.ι F (op j)).unop := by + simp [limitUnopIsoUnopColimit] + +@[reassoc (attr := simp)] +lemma limitUnopIsoUnopColimit_hom_comp_ι (F : Jᵒᵖ ⥤ Cᵒᵖ) [HasColimit F] (j : Jᵒᵖ) : + (limitUnopIsoUnopColimit F).hom ≫ (colimit.ι F j).unop = limit.π F.unop j.unop := by + simp [← Iso.eq_inv_comp] + /-- If `C` has colimits of shape `Jᵒᵖ`, we can construct limits in `Cᵒᵖ` of shape `J`. -/ theorem hasLimitsOfShape_op_of_hasColimitsOfShape [HasColimitsOfShape Jᵒᵖ C] : @@ -354,6 +414,66 @@ instance hasColimit_unop_of_hasLimit (F : Jᵒᵖ ⥤ Cᵒᵖ) [HasLimit F] : Ha { cocone := coconeUnopOfCone (limit.cone F) isColimit := isColimitCoconeUnopOfCone _ (limit.isLimit _) } +/-- The colimit of `F.op` is the opposite of `limit F`. -/ +def colimitOpIsoOpLimit (F : J ⥤ C) [HasLimit F] : + colimit F.op ≅ op (limit F) := + colimit.isoColimitCocone ⟨_, (limit.isLimit _).op⟩ + +@[reassoc (attr := simp)] +lemma ι_comp_colimitOpIsoOpLimit_hom (F : J ⥤ C) [HasLimit F] (j : Jᵒᵖ) : + colimit.ι F.op j ≫ (colimitOpIsoOpLimit F).hom = (limit.π F j.unop).op := by + simp [colimitOpIsoOpLimit] + +@[reassoc (attr := simp)] +lemma π_comp_colimitOpIsoOpLimit_inv (F : J ⥤ C) [HasLimit F] (j : J) : + (limit.π F j).op ≫ (colimitOpIsoOpLimit F).inv = colimit.ι F.op (op j) := by + simp [Iso.comp_inv_eq] + +/-- The colimit of `F.leftOp` is the unopposite of `limit F`. -/ +def colimitLeftOpIsoUnopLimit (F : J ⥤ Cᵒᵖ) [HasLimit F] : + colimit F.leftOp ≅ unop (limit F) := + colimit.isoColimitCocone ⟨_, isColimitCoconeLeftOpOfCone _ (limit.isLimit _)⟩ + +@[reassoc (attr := simp)] +lemma ι_comp_colimitLeftOpIsoUnopLimit_hom (F : J ⥤ Cᵒᵖ) [HasLimit F] (j : Jᵒᵖ) : + colimit.ι F.leftOp j ≫ (colimitLeftOpIsoUnopLimit F).hom = (limit.π F j.unop).unop := by + simp [colimitLeftOpIsoUnopLimit] + +@[reassoc (attr := simp)] +lemma π_comp_colimitLeftOpIsoUnopLimit_inv (F : J ⥤ Cᵒᵖ) [HasLimit F] (j : J) : + (limit.π F j).unop ≫ (colimitLeftOpIsoUnopLimit F).inv = colimit.ι F.leftOp (op j) := by + simp [Iso.comp_inv_eq] + +/-- The colimit of `F.rightOp` is the opposite of `limit F`. -/ +def colimitRightOpIsoUnopLimit (F : Jᵒᵖ ⥤ C) [HasLimit F] : + colimit F.rightOp ≅ op (limit F) := + colimit.isoColimitCocone ⟨_, isColimitCoconeRightOpOfCone _ (limit.isLimit _)⟩ + +@[reassoc (attr := simp)] +lemma ι_comp_colimitRightOpIsoUnopLimit_hom (F : Jᵒᵖ ⥤ C) [HasLimit F] (j : J) : + colimit.ι F.rightOp j ≫ (colimitRightOpIsoUnopLimit F).hom = (limit.π F (op j)).op := by + simp [colimitRightOpIsoUnopLimit] + +@[reassoc (attr := simp)] +lemma π_comp_colimitRightOpIsoUnopLimit_inv (F : Jᵒᵖ ⥤ C) [HasLimit F] (j : Jᵒᵖ) : + (limit.π F j).op ≫ (colimitRightOpIsoUnopLimit F).inv = colimit.ι F.rightOp j.unop := by + simp [Iso.comp_inv_eq] + +/-- The colimit of `F.unop` is the unopposite of `limit F`. -/ +def colimitUnopIsoOpLimit (F : Jᵒᵖ ⥤ Cᵒᵖ) [HasLimit F] : + colimit F.unop ≅ unop (limit F) := + colimit.isoColimitCocone ⟨_, isColimitCoconeUnopOfCone _ (limit.isLimit _)⟩ + +@[reassoc (attr := simp)] +lemma ι_comp_colimitUnopIsoOpLimit_hom (F : Jᵒᵖ ⥤ Cᵒᵖ) [HasLimit F] (j : J) : + colimit.ι F.unop j ≫ (colimitUnopIsoOpLimit F).hom = (limit.π F (op j)).unop := by + simp [colimitUnopIsoOpLimit] + +@[reassoc (attr := simp)] +lemma π_comp_colimitUnopIsoOpLimit_inv (F : Jᵒᵖ ⥤ Cᵒᵖ) [HasLimit F] (j : Jᵒᵖ) : + (limit.π F j).unop ≫ (colimitUnopIsoOpLimit F).inv = colimit.ι F.unop j.unop := by + simp [Iso.comp_inv_eq] + /-- If `C` has colimits of shape `Jᵒᵖ`, we can construct limits in `Cᵒᵖ` of shape `J`. -/ instance hasColimitsOfShape_op_of_hasLimitsOfShape [HasLimitsOfShape Jᵒᵖ C] : From 5ebc4a6a8356b57bef5e7a7d9f8e08c485424865 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Mon, 18 Nov 2024 14:56:38 +0000 Subject: [PATCH 071/829] feat(Data/seq): basic lemmas about Stream'.seq and the definition of length (#18589) --- Mathlib/Data/Seq/Seq.lean | 161 +++++++++++++++++++++++++++++++++++++- 1 file changed, 160 insertions(+), 1 deletion(-) diff --git a/Mathlib/Data/Seq/Seq.lean b/Mathlib/Data/Seq/Seq.lean index c781220c19e0c..76e8c653427c3 100644 --- a/Mathlib/Data/Seq/Seq.lean +++ b/Mathlib/Data/Seq/Seq.lean @@ -394,6 +394,9 @@ theorem ofList_get (l : List α) (n : ℕ) : (ofList l).get? n = l.get? n := theorem ofList_cons (a : α) (l : List α) : ofList (a::l) = cons a (ofList l) := by ext1 (_ | n) <;> rfl +theorem ofList_injective : Function.Injective (ofList : List α → _) := + fun _ _ h => List.ext_get? fun _ => congr_fun (Subtype.ext_iff.1 h) _ + /-- Embed an infinite stream as a sequence -/ @[coe] def ofStream (s : Stream' α) : Seq α := @@ -543,9 +546,13 @@ theorem get?_enum (s : Seq α) (n : ℕ) : get? (enum s) n = Option.map (Prod.mk theorem enum_nil : enum (nil : Seq α) = nil := rfl +/-- The length of a terminating sequence. -/ +def length (s : Seq α) (h : s.Terminates) : ℕ := + Nat.find h + /-- Convert a sequence which is known to terminate into a list -/ def toList (s : Seq α) (h : s.Terminates) : List α := - take (Nat.find h) s + take (length s h) s /-- Convert a sequence which is known not to terminate into a stream -/ def toStream (s : Seq α) (h : ¬s.Terminates) : Stream' α := fun n => @@ -568,6 +575,140 @@ theorem nil_append (s : Seq α) : append nil s = s := by dsimp exact ⟨rfl, s, rfl, rfl⟩ +@[simp] +theorem getElem?_take : ∀ (n k : ℕ) (s : Seq α), + (s.take k)[n]? = if n < k then s.get? n else none + | n, 0, s => by simp [take] + | n, k+1, s => by + rw [take] + cases h : destruct s with + | none => + simp [destruct_eq_nil h] + | some a => + match a with + | (x, r) => + rw [destruct_eq_cons h] + match n with + | 0 => simp + | n+1 => simp [List.get?_cons_succ, Nat.add_lt_add_iff_right, get?_cons_succ, getElem?_take] + +theorem terminatedAt_ofList (l : List α) : + (ofList l).TerminatedAt l.length := by + simp [ofList, TerminatedAt] + +theorem terminates_ofList (l : List α) : (ofList l).Terminates := + ⟨_, terminatedAt_ofList l⟩ + +theorem terminatedAt_nil : TerminatedAt (nil : Seq α) 0 := rfl + +@[simp] +theorem terminates_nil : Terminates (nil : Seq α) := ⟨0, rfl⟩ + +@[simp] +theorem length_nil : length (nil : Seq α) terminates_nil = 0 := rfl + +@[simp] +theorem get?_zero_eq_none {s : Seq α} : s.get? 0 = none ↔ s = nil := by + refine ⟨fun h => ?_, fun h => h ▸ rfl⟩ + ext1 n + exact le_stable s (Nat.zero_le _) h + +@[simp] theorem length_eq_zero {s : Seq α} {h : s.Terminates} : + s.length h = 0 ↔ s = nil := by + simp [length, TerminatedAt] + +theorem terminatedAt_zero_iff {s : Seq α} : s.TerminatedAt 0 ↔ s = nil := by + refine ⟨?_, ?_⟩ + · intro h + ext n + rw [le_stable _ (Nat.zero_le _) h] + simp + · rintro rfl + simp [TerminatedAt] + +/-- The statement of `length_le_iff'` does not assume that the sequence terminates. For a +simpler statement of the theorem where the sequence is known to terminate see `length_le_iff` -/ +theorem length_le_iff' {s : Seq α} {n : ℕ} : + (∃ h, s.length h ≤ n) ↔ s.TerminatedAt n := by + simp only [length, Nat.find_le_iff, TerminatedAt, Terminates, exists_prop] + refine ⟨?_, ?_⟩ + · rintro ⟨_, k, hkn, hk⟩ + exact le_stable s hkn hk + · intro hn + exact ⟨⟨n, hn⟩, ⟨n, le_rfl, hn⟩⟩ + +/-- The statement of `length_le_iff` assumes that the sequence terminates. For a +statement of the where the sequence is not known to terminate see `length_le_iff'` -/ +theorem length_le_iff {s : Seq α} {n : ℕ} {h : s.Terminates} : + s.length h ≤ n ↔ s.TerminatedAt n := by + rw [← length_le_iff']; simp [h] + +/-- The statement of `lt_length_iff'` does not assume that the sequence terminates. For a +simpler statement of the theorem where the sequence is known to terminate see `lt_length_iff` -/ +theorem lt_length_iff' {s : Seq α} {n : ℕ} : + (∀ h : s.Terminates, n < s.length h) ↔ ∃ a, a ∈ s.get? n := by + simp only [Terminates, TerminatedAt, length, Nat.lt_find_iff, forall_exists_index, Option.mem_def, + ← Option.ne_none_iff_exists', ne_eq] + refine ⟨?_, ?_⟩ + · intro h hn + exact h n hn n le_rfl hn + · intro hn _ _ k hkn hk + exact hn <| le_stable s hkn hk + +/-- The statement of `length_le_iff` assumes that the sequence terminates. For a +statement of the where the sequence is not known to terminate see `length_le_iff'` -/ +theorem lt_length_iff {s : Seq α} {n : ℕ} {h : s.Terminates} : + n < s.length h ↔ ∃ a, a ∈ s.get? n := by + rw [← lt_length_iff']; simp [h] + +theorem length_take_of_le_length {s : Seq α} {n : ℕ} + (hle : ∀ h : s.Terminates, n ≤ s.length h) : (s.take n).length = n := by + induction n generalizing s with + | zero => simp [take] + | succ n ih => + rw [take, destruct] + let ⟨a, ha⟩ := lt_length_iff'.1 (fun ht => lt_of_lt_of_le (Nat.succ_pos _) (hle ht)) + simp [Option.mem_def.1 ha] + rw [ih] + intro h + simp only [length, tail, Nat.le_find_iff, TerminatedAt, get?_mk, Stream'.tail] + intro m hmn hs + have := lt_length_iff'.1 (fun ht => (Nat.lt_of_succ_le (hle ht))) + rw [le_stable s (Nat.succ_le_of_lt hmn) hs] at this + simp at this + +@[simp] +theorem length_toList (s : Seq α) (h : s.Terminates) : (toList s h).length = length s h := by + rw [toList, length_take_of_le_length] + intro _ + exact le_rfl + +@[simp] +theorem getElem?_toList (s : Seq α) (h : s.Terminates) (n : ℕ) : (toList s h)[n]? = s.get? n := by + ext k + simp only [ofList, toList, get?_mk, Option.mem_def, getElem?_take, Nat.lt_find_iff, length, + Option.ite_none_right_eq_some, and_iff_right_iff_imp, TerminatedAt, List.get?_eq_getElem?] + intro h m hmn + let ⟨a, ha⟩ := ge_stable s hmn h + simp [ha] + +@[simp] +theorem ofList_toList (s : Seq α) (h : s.Terminates) : + ofList (toList s h) = s := by + ext n; simp [ofList, List.get?_eq_getElem?] + +@[simp] +theorem toList_ofList (l : List α) : toList (ofList l) (terminates_ofList l) = l := + ofList_injective (by simp) + +@[simp] +theorem toList_nil : toList (nil : Seq α) ⟨0, terminatedAt_zero_iff.2 rfl⟩ = [] := by + ext; simp [nil, toList, const] + +theorem getLast?_toList (s : Seq α) (h : s.Terminates) : + (toList s h).getLast? = s.get? (s.length h - 1) := by + rw [List.getLast?_eq_getElem?, getElem?_toList, length_toList] + @[simp] theorem cons_append (a : α) (s t) : append (cons a s) t = cons a (append s t) := destruct_eq_cons <| by @@ -646,6 +787,24 @@ theorem map_append (f : α → β) (s t) : map f (append s t) = append (map f s) theorem map_get? (f : α → β) : ∀ s n, get? (map f s) n = (get? s n).map f | ⟨_, _⟩, _ => rfl +@[simp] +theorem terminatedAt_map_iff {f : α → β} {s : Seq α} {n : ℕ} : + (map f s).TerminatedAt n ↔ s.TerminatedAt n := by + simp [TerminatedAt] + +@[simp] +theorem terminates_map_iff {f : α → β} {s : Seq α} : + (map f s).Terminates ↔ s.Terminates := by + simp [Terminates] + +@[simp] +theorem length_map {s : Seq α} {f : α → β} (h : (s.map f).Terminates) : + (s.map f).length h = s.length (terminates_map_iff.1 h) := by + rw [length] + congr + ext + simp + instance : Functor Seq where map := @map instance : LawfulFunctor Seq where From 7b7913144f3486bbb32583fa2e3925242452c7bf Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 18 Nov 2024 14:56:40 +0000 Subject: [PATCH 072/829] fix: do not count un-deprecations in `Mathlib/Deprecated` (#19184) [Reported on Zulip](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/Technical.20Debt.20Counters/near/482926966) --- scripts/technical-debt-metrics.sh | 12 +++++++++++- 1 file changed, 11 insertions(+), 1 deletion(-) diff --git a/scripts/technical-debt-metrics.sh b/scripts/technical-debt-metrics.sh index e607f9569cd79..eb24ded1ce153 100755 --- a/scripts/technical-debt-metrics.sh +++ b/scripts/technical-debt-metrics.sh @@ -79,7 +79,6 @@ titlesPathsAndRegexes=( "skipAssignedInstances flags" "*" "set_option tactic.skipAssignedInstances" "adaptation notes" "*" "adaptation_note" "disabled simpNF lints" "*" "nolint simpNF" - "disabled deprecation lints" "*" "set_option linter.deprecated false" "erw" "*" "erw \[" "maxHeartBeats modifications" ":^MathlibTest" "^ *set_option .*maxHeartbeats" ) @@ -98,6 +97,17 @@ for i in ${!titlesPathsAndRegexes[@]}; do fi done +# count total number of `set_option linter.deprecated false` outside of `Mathlib/Deprecated` +deprecs="$(git grep -c -- "set_option linter.deprecated false" -- ":^Mathlib/Deprecated" | + awk -F: 'BEGIN{total=0} {total+=$2} END{print total}')" + +# count the `linter.deprecated` exceptions that are themselves followed by `deprecated ...(since` +# we subtract these from `deprecs` +doubleDeprecs="$(git grep -A1 -- "set_option linter.deprecated false" -- ":^Mathlib/Deprecated" | + grep -c "deprecated .*(since")" + +printf '%s|disabled deprecation lints\n' "$(( deprecs - doubleDeprecs ))" + printf '%s|%s\n' "$(grep -c 'docBlame' scripts/nolints.json)" "documentation nolint entries" # We count the number of large files, making sure to avoid counting the test file `MathlibTest/Lint.lean`. printf '%s|%s\n' "$(git grep '^set_option linter.style.longFile [0-9]*' Mathlib | wc -l)" "large files" From d4a83c927887150c0a11c4102350ed13c27f2660 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Mon, 18 Nov 2024 15:16:39 +0000 Subject: [PATCH 073/829] chore: add some simp and fun_prop attributes (#18874) * Also add a test file that uses the `simp` lemmas. * Add/move the `fun_prop` lemmas to the compositional lemmas. [Zulip](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/Derivative.20cannot.20simp) --- Mathlib/Analysis/Calculus/Deriv/Basic.lean | 2 +- .../Analysis/SpecialFunctions/ExpDeriv.lean | 12 ++-- .../SpecialFunctions/Trigonometric/Deriv.lean | 64 +++++++++---------- MathlibTest/Deriv.lean | 25 ++++++++ 4 files changed, 66 insertions(+), 37 deletions(-) create mode 100644 MathlibTest/Deriv.lean diff --git a/Mathlib/Analysis/Calculus/Deriv/Basic.lean b/Mathlib/Analysis/Calculus/Deriv/Basic.lean index ce5715d4dccc9..2980e8e408490 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Basic.lean @@ -68,7 +68,7 @@ and they more frequently lead to the desired result. We set up the simplifier so that it can compute the derivative of simple functions. For instance, ```lean example (x : ℝ) : - deriv (fun x ↦ cos (sin x) * exp x) x = (cos(sin(x))-sin(sin(x))*cos(x))*exp(x) := by + deriv (fun x ↦ cos (sin x) * exp x) x = (cos (sin x) - sin (sin x) * cos x) * exp x := by simp; ring ``` diff --git a/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean b/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean index 8506d9e055ab1..3a485d339830b 100644 --- a/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean @@ -73,9 +73,11 @@ theorem hasDerivAt_exp (x : ℂ) : HasDerivAt exp (exp x) x := by simp only [Metric.mem_ball, dist_zero_right, norm_pow] exact fun z hz => exp_bound_sq x z hz.le +@[simp] theorem differentiable_exp : Differentiable 𝕜 exp := fun x => (hasDerivAt_exp x).differentiableAt.restrictScalars 𝕜 +@[simp] theorem differentiableAt_exp {x : ℂ} : DifferentiableAt 𝕜 exp x := differentiable_exp x @@ -148,7 +150,7 @@ theorem DifferentiableWithinAt.cexp (hf : DifferentiableWithinAt 𝕜 f s x) : DifferentiableWithinAt 𝕜 (fun x => Complex.exp (f x)) s x := hf.hasFDerivWithinAt.cexp.differentiableWithinAt -@[simp] +@[simp, fun_prop] theorem DifferentiableAt.cexp (hc : DifferentiableAt 𝕜 f x) : DifferentiableAt 𝕜 (fun x => Complex.exp (f x)) x := hc.hasFDerivAt.cexp.differentiableAt @@ -156,7 +158,7 @@ theorem DifferentiableAt.cexp (hc : DifferentiableAt 𝕜 f x) : theorem DifferentiableOn.cexp (hc : DifferentiableOn 𝕜 f s) : DifferentiableOn 𝕜 (fun x => Complex.exp (f x)) s := fun x h => (hc x h).cexp -@[simp] +@[simp, fun_prop] theorem Differentiable.cexp (hc : Differentiable 𝕜 f) : Differentiable 𝕜 fun x => Complex.exp (f x) := fun x => (hc x).cexp @@ -231,8 +233,10 @@ theorem hasDerivAt_exp (x : ℝ) : HasDerivAt exp (exp x) x := theorem contDiff_exp {n : ℕ∞} : ContDiff ℝ n exp := Complex.contDiff_exp.real_of_complex +@[simp] theorem differentiable_exp : Differentiable ℝ exp := fun x => (hasDerivAt_exp x).differentiableAt +@[simp] theorem differentiableAt_exp {x : ℝ} : DifferentiableAt ℝ exp x := differentiable_exp x @@ -316,7 +320,7 @@ theorem DifferentiableWithinAt.exp (hf : DifferentiableWithinAt ℝ f s x) : DifferentiableWithinAt ℝ (fun x => Real.exp (f x)) s x := hf.hasFDerivWithinAt.exp.differentiableWithinAt -@[simp] +@[simp, fun_prop] theorem DifferentiableAt.exp (hc : DifferentiableAt ℝ f x) : DifferentiableAt ℝ (fun x => Real.exp (f x)) x := hc.hasFDerivAt.exp.differentiableAt @@ -324,7 +328,7 @@ theorem DifferentiableAt.exp (hc : DifferentiableAt ℝ f x) : theorem DifferentiableOn.exp (hc : DifferentiableOn ℝ f s) : DifferentiableOn ℝ (fun x => Real.exp (f x)) s := fun x h => (hc x h).exp -@[simp] +@[simp, fun_prop] theorem Differentiable.exp (hc : Differentiable ℝ f) : Differentiable ℝ fun x => Real.exp (f x) := fun x => (hc x).exp diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Deriv.lean index 862ffd739dff4..0a5c2533c4cd0 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Deriv.lean @@ -45,10 +45,10 @@ theorem contDiff_sin {n} : ContDiff ℂ n sin := (((contDiff_neg.mul contDiff_const).cexp.sub (contDiff_id.mul contDiff_const).cexp).mul contDiff_const).div_const _ -@[fun_prop] +@[simp] theorem differentiable_sin : Differentiable ℂ sin := fun x => (hasDerivAt_sin x).differentiableAt -@[fun_prop] +@[simp] theorem differentiableAt_sin {x : ℂ} : DifferentiableAt ℂ sin x := differentiable_sin x @@ -72,10 +72,10 @@ theorem hasDerivAt_cos (x : ℂ) : HasDerivAt cos (-sin x) x := theorem contDiff_cos {n} : ContDiff ℂ n cos := ((contDiff_id.mul contDiff_const).cexp.add (contDiff_neg.mul contDiff_const).cexp).div_const _ -@[fun_prop] +@[simp] theorem differentiable_cos : Differentiable ℂ cos := fun x => (hasDerivAt_cos x).differentiableAt -@[fun_prop] +@[simp] theorem differentiableAt_cos {x : ℂ} : DifferentiableAt ℂ cos x := differentiable_cos x @@ -102,10 +102,10 @@ theorem hasDerivAt_sinh (x : ℂ) : HasDerivAt sinh (cosh x) x := theorem contDiff_sinh {n} : ContDiff ℂ n sinh := (contDiff_exp.sub contDiff_neg.cexp).div_const _ -@[fun_prop] +@[simp] theorem differentiable_sinh : Differentiable ℂ sinh := fun x => (hasDerivAt_sinh x).differentiableAt -@[fun_prop] +@[simp] theorem differentiableAt_sinh {x : ℂ} : DifferentiableAt ℂ sinh x := differentiable_sinh x @@ -129,10 +129,10 @@ theorem hasDerivAt_cosh (x : ℂ) : HasDerivAt cosh (sinh x) x := theorem contDiff_cosh {n} : ContDiff ℂ n cosh := (contDiff_exp.add contDiff_neg.cexp).div_const _ -@[fun_prop] +@[simp] theorem differentiable_cosh : Differentiable ℂ cosh := fun x => (hasDerivAt_cosh x).differentiableAt -@[fun_prop] +@[simp] theorem differentiableAt_cosh {x : ℂ} : DifferentiableAt ℂ cosh x := differentiable_cosh x @@ -274,7 +274,7 @@ theorem DifferentiableWithinAt.ccos (hf : DifferentiableWithinAt ℂ f s x) : DifferentiableWithinAt ℂ (fun x => Complex.cos (f x)) s x := hf.hasFDerivWithinAt.ccos.differentiableWithinAt -@[simp] +@[simp, fun_prop] theorem DifferentiableAt.ccos (hc : DifferentiableAt ℂ f x) : DifferentiableAt ℂ (fun x => Complex.cos (f x)) x := hc.hasFDerivAt.ccos.differentiableAt @@ -282,7 +282,7 @@ theorem DifferentiableAt.ccos (hc : DifferentiableAt ℂ f x) : theorem DifferentiableOn.ccos (hc : DifferentiableOn ℂ f s) : DifferentiableOn ℂ (fun x => Complex.cos (f x)) s := fun x h => (hc x h).ccos -@[simp] +@[simp, fun_prop] theorem Differentiable.ccos (hc : Differentiable ℂ f) : Differentiable ℂ fun x => Complex.cos (f x) := fun x => (hc x).ccos @@ -329,7 +329,7 @@ theorem DifferentiableWithinAt.csin (hf : DifferentiableWithinAt ℂ f s x) : DifferentiableWithinAt ℂ (fun x => Complex.sin (f x)) s x := hf.hasFDerivWithinAt.csin.differentiableWithinAt -@[simp] +@[simp, fun_prop] theorem DifferentiableAt.csin (hc : DifferentiableAt ℂ f x) : DifferentiableAt ℂ (fun x => Complex.sin (f x)) x := hc.hasFDerivAt.csin.differentiableAt @@ -337,7 +337,7 @@ theorem DifferentiableAt.csin (hc : DifferentiableAt ℂ f x) : theorem DifferentiableOn.csin (hc : DifferentiableOn ℂ f s) : DifferentiableOn ℂ (fun x => Complex.sin (f x)) s := fun x h => (hc x h).csin -@[simp] +@[simp, fun_prop] theorem Differentiable.csin (hc : Differentiable ℂ f) : Differentiable ℂ fun x => Complex.sin (f x) := fun x => (hc x).csin @@ -384,7 +384,7 @@ theorem DifferentiableWithinAt.ccosh (hf : DifferentiableWithinAt ℂ f s x) : DifferentiableWithinAt ℂ (fun x => Complex.cosh (f x)) s x := hf.hasFDerivWithinAt.ccosh.differentiableWithinAt -@[simp] +@[simp, fun_prop] theorem DifferentiableAt.ccosh (hc : DifferentiableAt ℂ f x) : DifferentiableAt ℂ (fun x => Complex.cosh (f x)) x := hc.hasFDerivAt.ccosh.differentiableAt @@ -392,7 +392,7 @@ theorem DifferentiableAt.ccosh (hc : DifferentiableAt ℂ f x) : theorem DifferentiableOn.ccosh (hc : DifferentiableOn ℂ f s) : DifferentiableOn ℂ (fun x => Complex.cosh (f x)) s := fun x h => (hc x h).ccosh -@[simp] +@[simp, fun_prop] theorem Differentiable.ccosh (hc : Differentiable ℂ f) : Differentiable ℂ fun x => Complex.cosh (f x) := fun x => (hc x).ccosh @@ -439,7 +439,7 @@ theorem DifferentiableWithinAt.csinh (hf : DifferentiableWithinAt ℂ f s x) : DifferentiableWithinAt ℂ (fun x => Complex.sinh (f x)) s x := hf.hasFDerivWithinAt.csinh.differentiableWithinAt -@[simp] +@[simp, fun_prop] theorem DifferentiableAt.csinh (hc : DifferentiableAt ℂ f x) : DifferentiableAt ℂ (fun x => Complex.sinh (f x)) x := hc.hasFDerivAt.csinh.differentiableAt @@ -447,7 +447,7 @@ theorem DifferentiableAt.csinh (hc : DifferentiableAt ℂ f x) : theorem DifferentiableOn.csinh (hc : DifferentiableOn ℂ f s) : DifferentiableOn ℂ (fun x => Complex.sinh (f x)) s := fun x h => (hc x h).csinh -@[simp] +@[simp, fun_prop] theorem Differentiable.csinh (hc : Differentiable ℂ f) : Differentiable ℂ fun x => Complex.sinh (f x) := fun x => (hc x).csinh @@ -490,10 +490,10 @@ theorem hasDerivAt_sin (x : ℝ) : HasDerivAt sin (cos x) x := theorem contDiff_sin {n} : ContDiff ℝ n sin := Complex.contDiff_sin.real_of_complex -@[fun_prop] +@[simp] theorem differentiable_sin : Differentiable ℝ sin := fun x => (hasDerivAt_sin x).differentiableAt -@[fun_prop] +@[simp] theorem differentiableAt_sin : DifferentiableAt ℝ sin x := differentiable_sin x @@ -510,10 +510,10 @@ theorem hasDerivAt_cos (x : ℝ) : HasDerivAt cos (-sin x) x := theorem contDiff_cos {n} : ContDiff ℝ n cos := Complex.contDiff_cos.real_of_complex -@[fun_prop] +@[simp] theorem differentiable_cos : Differentiable ℝ cos := fun x => (hasDerivAt_cos x).differentiableAt -@[fun_prop] +@[simp] theorem differentiableAt_cos : DifferentiableAt ℝ cos x := differentiable_cos x @@ -533,10 +533,10 @@ theorem hasDerivAt_sinh (x : ℝ) : HasDerivAt sinh (cosh x) x := theorem contDiff_sinh {n} : ContDiff ℝ n sinh := Complex.contDiff_sinh.real_of_complex -@[fun_prop] +@[simp] theorem differentiable_sinh : Differentiable ℝ sinh := fun x => (hasDerivAt_sinh x).differentiableAt -@[fun_prop] +@[simp] theorem differentiableAt_sinh : DifferentiableAt ℝ sinh x := differentiable_sinh x @@ -553,10 +553,10 @@ theorem hasDerivAt_cosh (x : ℝ) : HasDerivAt cosh (sinh x) x := theorem contDiff_cosh {n} : ContDiff ℝ n cosh := Complex.contDiff_cosh.real_of_complex -@[fun_prop] +@[simp] theorem differentiable_cosh : Differentiable ℝ cosh := fun x => (hasDerivAt_cosh x).differentiableAt -@[fun_prop] +@[simp] theorem differentiableAt_cosh : DifferentiableAt ℝ cosh x := differentiable_cosh x @@ -786,7 +786,7 @@ theorem DifferentiableWithinAt.cos (hf : DifferentiableWithinAt ℝ f s x) : DifferentiableWithinAt ℝ (fun x => Real.cos (f x)) s x := hf.hasFDerivWithinAt.cos.differentiableWithinAt -@[simp] +@[simp, fun_prop] theorem DifferentiableAt.cos (hc : DifferentiableAt ℝ f x) : DifferentiableAt ℝ (fun x => Real.cos (f x)) x := hc.hasFDerivAt.cos.differentiableAt @@ -794,7 +794,7 @@ theorem DifferentiableAt.cos (hc : DifferentiableAt ℝ f x) : theorem DifferentiableOn.cos (hc : DifferentiableOn ℝ f s) : DifferentiableOn ℝ (fun x => Real.cos (f x)) s := fun x h => (hc x h).cos -@[simp] +@[simp, fun_prop] theorem Differentiable.cos (hc : Differentiable ℝ f) : Differentiable ℝ fun x => Real.cos (f x) := fun x => (hc x).cos @@ -839,7 +839,7 @@ theorem DifferentiableWithinAt.sin (hf : DifferentiableWithinAt ℝ f s x) : DifferentiableWithinAt ℝ (fun x => Real.sin (f x)) s x := hf.hasFDerivWithinAt.sin.differentiableWithinAt -@[simp] +@[simp, fun_prop] theorem DifferentiableAt.sin (hc : DifferentiableAt ℝ f x) : DifferentiableAt ℝ (fun x => Real.sin (f x)) x := hc.hasFDerivAt.sin.differentiableAt @@ -847,7 +847,7 @@ theorem DifferentiableAt.sin (hc : DifferentiableAt ℝ f x) : theorem DifferentiableOn.sin (hc : DifferentiableOn ℝ f s) : DifferentiableOn ℝ (fun x => Real.sin (f x)) s := fun x h => (hc x h).sin -@[simp] +@[simp, fun_prop] theorem Differentiable.sin (hc : Differentiable ℝ f) : Differentiable ℝ fun x => Real.sin (f x) := fun x => (hc x).sin @@ -892,7 +892,7 @@ theorem DifferentiableWithinAt.cosh (hf : DifferentiableWithinAt ℝ f s x) : DifferentiableWithinAt ℝ (fun x => Real.cosh (f x)) s x := hf.hasFDerivWithinAt.cosh.differentiableWithinAt -@[simp] +@[simp, fun_prop] theorem DifferentiableAt.cosh (hc : DifferentiableAt ℝ f x) : DifferentiableAt ℝ (fun x => Real.cosh (f x)) x := hc.hasFDerivAt.cosh.differentiableAt @@ -900,7 +900,7 @@ theorem DifferentiableAt.cosh (hc : DifferentiableAt ℝ f x) : theorem DifferentiableOn.cosh (hc : DifferentiableOn ℝ f s) : DifferentiableOn ℝ (fun x => Real.cosh (f x)) s := fun x h => (hc x h).cosh -@[simp] +@[simp, fun_prop] theorem Differentiable.cosh (hc : Differentiable ℝ f) : Differentiable ℝ fun x => Real.cosh (f x) := fun x => (hc x).cosh @@ -947,7 +947,7 @@ theorem DifferentiableWithinAt.sinh (hf : DifferentiableWithinAt ℝ f s x) : DifferentiableWithinAt ℝ (fun x => Real.sinh (f x)) s x := hf.hasFDerivWithinAt.sinh.differentiableWithinAt -@[simp] +@[simp, fun_prop] theorem DifferentiableAt.sinh (hc : DifferentiableAt ℝ f x) : DifferentiableAt ℝ (fun x => Real.sinh (f x)) x := hc.hasFDerivAt.sinh.differentiableAt @@ -955,7 +955,7 @@ theorem DifferentiableAt.sinh (hc : DifferentiableAt ℝ f x) : theorem DifferentiableOn.sinh (hc : DifferentiableOn ℝ f s) : DifferentiableOn ℝ (fun x => Real.sinh (f x)) s := fun x h => (hc x h).sinh -@[simp] +@[simp, fun_prop] theorem Differentiable.sinh (hc : Differentiable ℝ f) : Differentiable ℝ fun x => Real.sinh (f x) := fun x => (hc x).sinh diff --git a/MathlibTest/Deriv.lean b/MathlibTest/Deriv.lean new file mode 100644 index 0000000000000..7c45e91681b70 --- /dev/null +++ b/MathlibTest/Deriv.lean @@ -0,0 +1,25 @@ +import Mathlib + +/-! Test that `simp` can prove some lemmas about derivatives. -/ + +open Real + +example (x : ℝ) : deriv (fun x => cos x + 2 * sin x) x = -sin x + 2 * cos x := by + simp + +example (x : ℝ) : + deriv (fun x ↦ cos (sin x) * exp x) x = (cos (sin x) - sin (sin x) * cos x) * exp x := by + simp; ring + +/- for more complicated examples (with more nested functions) you need to increase the +`maxDischargeDepth`. -/ + +example (x : ℝ) : + deriv (fun x ↦ sin (sin (sin x)) + sin x) x = + cos (sin (sin x)) * (cos (sin x) * cos x) + cos x := by + simp (maxDischargeDepth := 3) + +example (x : ℝ) : + deriv (fun x ↦ sin (sin (sin x)) ^ 10 + sin x) x = + 10 * sin (sin (sin x)) ^ 9 * (cos (sin (sin x)) * (cos (sin x) * cos x)) + cos x := by + simp (maxDischargeDepth := 4) From bdfd6f2dedc09ca171e0a38fc456a338c336a40b Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 18 Nov 2024 15:35:41 +0000 Subject: [PATCH 074/829] feat: add spoiler with comment to maintainer merge (#19069) For this PR, it would produce a message like adomani requested a maintainer **delegate** from comment on PR [#19069](url):
feat: add spoiler with comment to maintainer merge This is a great addition, thanks! maintainer delegate Maintainers, please check that it passes CI!
--- .github/workflows/maintainer_merge.yml | 5 ++++- scripts/maintainer_merge_message.sh | 4 +++- 2 files changed, 7 insertions(+), 2 deletions(-) diff --git a/.github/workflows/maintainer_merge.yml b/.github/workflows/maintainer_merge.yml index be9e7ee77facb..87ab7583093b7 100644 --- a/.github/workflows/maintainer_merge.yml +++ b/.github/workflows/maintainer_merge.yml @@ -33,11 +33,14 @@ jobs: id: merge_or_delegate run: | COMMENT="${COMMENT_EVENT}${COMMENT_REVIEW}" + # we strip `\r` since line endings from GitHub contain this character COMMENT="${COMMENT//$'\r'/}" + # for debugging, we print some information printf '%s' "${COMMENT}" | hexdump -cC printf 'Comment:"%s"\n' "${COMMENT}" + m_or_d="$(printf '%s' "${COMMENT}" | sed -n 's=^maintainer *\(merge\|delegate\) *$=\1=p' | head -1)" @@ -77,7 +80,7 @@ jobs: echo "" message="$( - ./scripts/maintainer_merge_message.sh "${AUTHOR}" "${{ steps.merge_or_delegate.outputs.mOrD }}" "${EVENT_NAME}" "${PR_NUMBER}" "${PR_URL}" "${PR_TITLE_ISSUE}${PR_TITLE_PR}" + ./scripts/maintainer_merge_message.sh "${AUTHOR}" "${{ steps.merge_or_delegate.outputs.mOrD }}" "${EVENT_NAME}" "${PR_NUMBER}" "${PR_URL}" "${PR_TITLE_ISSUE}${PR_TITLE_PR}" "${COMMENT_EVENT}${COMMENT_REVIEW}" )" printf 'title<&2 echo "PR_URL: '${PR_URL}'" >&2 echo "title: '${PR_TITLE}'" >&2 echo "EVENT_NAME: '${EVENT_NAME}'" +>&2 printf 'COMMENT\n%s\nEND COMMENT\n' "${PR_COMMENT}" printf '%s requested a maintainer **%s** from %s on PR [#%s](%s):\n' "${AUTHOR}" "${M_or_D}" "${SOURCE}" "${PR}" "${URL}" -printf '> %s\n' "${PR_TITLE}" +printf '```spoiler %s\n%s\n```\n' "${PR_TITLE}" "${PR_COMMENT}" From f2ac82b1b3e131a7cc69451ffed99d54f52aa8a1 Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Mon, 18 Nov 2024 15:35:42 +0000 Subject: [PATCH 075/829] chore(FieldTheory/PurelyInseparable): fix typo `_of_isIntegral'` -> `_of_isSeparable'` (#19182) --- Mathlib/FieldTheory/PurelyInseparable.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/FieldTheory/PurelyInseparable.lean b/Mathlib/FieldTheory/PurelyInseparable.lean index f077e7c7ee4e4..6085b263e3ab1 100644 --- a/Mathlib/FieldTheory/PurelyInseparable.lean +++ b/Mathlib/FieldTheory/PurelyInseparable.lean @@ -836,7 +836,7 @@ theorem LinearIndependent.map_pow_expChar_pow_of_isSeparable [Algebra.IsSeparabl /-- If `E / F` is a field extension of exponential characteristic `q`, if `{ u_i }` is a family of separable elements of `E` which is `F`-linearly independent, then `{ u_i ^ (q ^ n) }` is also `F`-linearly independent for any natural number `n`. -/ -theorem LinearIndependent.map_pow_expChar_pow_of_isIntegral' +theorem LinearIndependent.map_pow_expChar_pow_of_isSeparable' (hsep : ∀ i : ι, IsSeparable F (v i)) (h : LinearIndependent F v) : LinearIndependent F (v · ^ q ^ n) := by let E' := adjoin F (Set.range v) @@ -990,7 +990,7 @@ theorem LinearIndependent.map_of_isPurelyInseparable_of_isSeparable [IsPurelyIns contrapose! refine fun hs ↦ (injective_iff_map_eq_zero _).mp (algebraMap F E).injective _ ?_ rw [hlF, Finsupp.not_mem_support_iff.1 hs, zero_pow this] - replace h := linearIndependent_iff.1 (h.map_pow_expChar_pow_of_isIntegral' q n hsep) lF₀ <| by + replace h := linearIndependent_iff.1 (h.map_pow_expChar_pow_of_isSeparable' q n hsep) lF₀ <| by replace hl := congr($hl ^ q ^ n) rw [linearCombination_apply, Finsupp.sum, sum_pow_char_pow, zero_pow this] at hl rw [← hl, linearCombination_apply, onFinset_sum _ (fun _ ↦ by exact zero_smul _ _)] From b04165d049454c0676bf3fa5e3633555af887f6e Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 18 Nov 2024 15:35:44 +0000 Subject: [PATCH 076/829] doc: `Result.isRat` assumes coprimality (#19207) Confirmed in dicussion with Mario; `norm_num` relies on this assumption when proving equality of rationals by `rfl`. --- Mathlib/Tactic/NormNum/Result.lean | 7 ++++--- 1 file changed, 4 insertions(+), 3 deletions(-) diff --git a/Mathlib/Tactic/NormNum/Result.lean b/Mathlib/Tactic/NormNum/Result.lean index a8a58dcf2cb7a..85a73f0db9eca 100644 --- a/Mathlib/Tactic/NormNum/Result.lean +++ b/Mathlib/Tactic/NormNum/Result.lean @@ -280,9 +280,10 @@ and `proof : isInt x (.negOfNat lit)`. -/ ∀ (inst : Q(Ring $α) := by assumption) (lit : Q(ℕ)) (proof : Q(IsInt $x (.negOfNat $lit))), Result x := Result'.isNegNat -/-- The result is `proof : isRat x n d`, where `n` is either `.ofNat lit` or `.negOfNat lit` -with `lit` a raw nat literal and `d` is a raw nat literal (not 0 or 1), -and `q` is the value of `n / d`. -/ +/-- The result is `proof : isRat x n d`, +where `n` is either `.ofNat lit` or `.negOfNat lit` with `lit` a raw nat literal, +`d` is a raw nat literal (not 0 or 1), +`n` and `d` are coprime, and `q` is the value of `n / d`. -/ @[match_pattern, inline] def Result.isRat {α : Q(Type u)} {x : Q($α)} : ∀ (inst : Q(DivisionRing $α) := by assumption) (q : Rat) (n : Q(ℤ)) (d : Q(ℕ)) (proof : Q(IsRat $x $n $d)), Result x := Result'.isRat From 72194f73ba222e484268e240b0ac989981fa71e8 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Mon, 18 Nov 2024 16:07:34 +0000 Subject: [PATCH 077/829] feat(Algebra): `Submodule R A` is algebra over `Ideal A` (#18493) Also upgrade `Submodule.mapHom` and `Ideal.mapHom` to RingHoms and show the former is also an `AlgHom` (`Submodule.mapAlgHom`). --- Mathlib/Algebra/Algebra/Operations.lean | 19 ++++++----- .../Algebra/Group/Submonoid/Pointwise.lean | 4 +++ Mathlib/RingTheory/Ideal/Maps.lean | 24 +++++++------ Mathlib/RingTheory/Ideal/Operations.lean | 34 ++++++++++++++++++- 4 files changed, 61 insertions(+), 20 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Operations.lean b/Mathlib/Algebra/Algebra/Operations.lean index 33686a95e9ef4..8decae0eb7f39 100644 --- a/Mathlib/Algebra/Algebra/Operations.lean +++ b/Mathlib/Algebra/Algebra/Operations.lean @@ -159,6 +159,9 @@ theorem mul_le_mul_left (h : M ≤ N) : M * P ≤ N * P := theorem mul_le_mul_right (h : N ≤ P) : M * N ≤ M * P := AddSubmonoid.mul_le_mul_right h +theorem mul_comm_of_commute (h : ∀ m ∈ M, ∀ n ∈ N, Commute m n) : M * N = N * M := + toAddSubmonoid_injective <| AddSubmonoid.mul_comm_of_commute h + variable (M N P) theorem mul_sup : M * (N ⊔ P) = M * N ⊔ M * P := @@ -505,14 +508,17 @@ protected theorem pow_induction_on_right {C : A → Prop} (hr : ∀ r : R, C (al (fun x y _i _hx _hy => hadd x y) (fun _i _x _hx => hmul _) hx -/-- `Submonoid.map` as a `MonoidWithZeroHom`, when applied to `AlgHom`s. -/ +/-- `Submonoid.map` as a `RingHom`, when applied to an `AlgHom`. -/ @[simps] def mapHom {A'} [Semiring A'] [Algebra R A'] (f : A →ₐ[R] A') : - Submodule R A →*₀ Submodule R A' where + Submodule R A →+* Submodule R A' where toFun := map f.toLinearMap map_zero' := Submodule.map_bot _ + map_add' := (Submodule.map_sup · · _) map_one' := Submodule.map_one _ - map_mul' _ _ := Submodule.map_mul _ _ _ + map_mul' := (Submodule.map_mul · · _) + +theorem mapHom_id : mapHom (.id R A) = .id _ := RingHom.ext map_id /-- The ring of submodules of the opposite algebra is isomorphic to the opposite ring of submodules. -/ @@ -553,14 +559,11 @@ theorem map_unop_pow (n : ℕ) (M : Submodule R Aᵐᵒᵖ) : on either side). -/ @[simps] def span.ringHom : SetSemiring A →+* Submodule R A where - -- Note: the hint `(α := A)` is new in https://github.com/leanprover-community/mathlib4/pull/8386 - toFun s := Submodule.span R (SetSemiring.down (α := A) s) + toFun s := Submodule.span R (SetSemiring.down s) map_zero' := span_empty map_one' := one_eq_span.symm map_add' := span_union - map_mul' s t := by - dsimp only -- Porting note: new, needed due to new-style structures - rw [SetSemiring.down_mul, span_mul_span] + map_mul' s t := by simp_rw [SetSemiring.down_mul, span_mul_span] section diff --git a/Mathlib/Algebra/Group/Submonoid/Pointwise.lean b/Mathlib/Algebra/Group/Submonoid/Pointwise.lean index 356b2543f3dde..2c5e53bea3b80 100644 --- a/Mathlib/Algebra/Group/Submonoid/Pointwise.lean +++ b/Mathlib/Algebra/Group/Submonoid/Pointwise.lean @@ -595,6 +595,10 @@ theorem iSup_mul (S : ι → AddSubmonoid R) (T : AddSubmonoid R) : (⨆ i, S i) theorem mul_iSup (T : AddSubmonoid R) (S : ι → AddSubmonoid R) : (T * ⨆ i, S i) = ⨆ i, T * S i := smul_iSup T S +theorem mul_comm_of_commute (h : ∀ m ∈ M, ∀ n ∈ N, Commute m n) : M * N = N * M := + le_antisymm (mul_le.mpr fun m hm n hn ↦ h m hm n hn ▸ mul_mem_mul hn hm) + (mul_le.mpr fun n hn m hm ↦ h m hm n hn ▸ mul_mem_mul hm hn) + end NonUnitalNonAssocSemiring section NonUnitalNonAssocRing diff --git a/Mathlib/RingTheory/Ideal/Maps.lean b/Mathlib/RingTheory/Ideal/Maps.lean index 48d4c2d1cea4c..7a67ded065aff 100644 --- a/Mathlib/RingTheory/Ideal/Maps.lean +++ b/Mathlib/RingTheory/Ideal/Maps.lean @@ -248,15 +248,16 @@ theorem smul_top_eq_map {R S : Type*} [CommSemiring R] [CommSemiring S] [Algebra congrArg _ <| Eq.trans (Ideal.smul_eq_mul _ _) (Ideal.mul_top _) @[simp] -theorem coe_restrictScalars {R S : Type*} [CommSemiring R] [Semiring S] [Algebra R S] - (I : Ideal S) : (I.restrictScalars R : Set S) = ↑I := +theorem coe_restrictScalars {R S : Type*} [Semiring R] [Semiring S] [Module R S] + [IsScalarTower R S S] (I : Ideal S) : (I.restrictScalars R : Set S) = ↑I := rfl /-- The smallest `S`-submodule that contains all `x ∈ I * y ∈ J` is also the smallest `R`-submodule that does so. -/ @[simp] -theorem restrictScalars_mul {R S : Type*} [CommSemiring R] [CommSemiring S] [Algebra R S] - (I J : Ideal S) : (I * J).restrictScalars R = I.restrictScalars R * J.restrictScalars R := +theorem restrictScalars_mul {R S : Type*} [Semiring R] [Semiring S] [Module R S] + [IsScalarTower R S S] (I J : Ideal S) : + (I * J).restrictScalars R = I.restrictScalars R * J.restrictScalars R := rfl section Surjective @@ -493,13 +494,13 @@ end Ring section CommRing -variable {F : Type*} [CommRing R] [CommRing S] +variable {F : Type*} [CommSemiring R] [CommSemiring S] variable [FunLike F R S] [rc : RingHomClass F R S] variable (f : F) -variable {I J : Ideal R} {K L : Ideal S} -variable (I J K L) +variable (I J : Ideal R) (K L : Ideal S) -theorem map_mul : map f (I * J) = map f I * map f J := +theorem map_mul {R} [Semiring R] [FunLike F R S] [RingHomClass F R S] (f : F) (I J : Ideal R) : + map f (I * J) = map f I * map f J := le_antisymm (map_le_iff_le_comap.2 <| mul_le.2 fun r hri s hsj => @@ -511,12 +512,13 @@ theorem map_mul : map f (I * J) = map f I * map f J := Set.singleton_subset_iff.2 <| hfri ▸ hfsj ▸ by rw [← _root_.map_mul]; exact mem_map_of_mem f (mul_mem_mul hri hsj))) -/-- The pushforward `Ideal.map` as a monoid-with-zero homomorphism. -/ +/-- The pushforward `Ideal.map` as a (semi)ring homomorphism. -/ @[simps] -def mapHom : Ideal R →*₀ Ideal S where +def mapHom : Ideal R →+* Ideal S where toFun := map f - map_mul' I J := Ideal.map_mul f I J + map_mul' := Ideal.map_mul f map_one' := by simp only [one_eq_top]; exact Ideal.map_top f + map_add' I J := Ideal.map_sup f I J map_zero' := Ideal.map_bot protected theorem map_pow (n : ℕ) : map f (I ^ n) = map f I ^ n := diff --git a/Mathlib/RingTheory/Ideal/Operations.lean b/Mathlib/RingTheory/Ideal/Operations.lean index 6456757fd98c5..e18052a266b00 100644 --- a/Mathlib/RingTheory/Ideal/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Operations.lean @@ -1241,7 +1241,6 @@ namespace Submodule variable {R : Type u} {M : Type v} variable [CommSemiring R] [AddCommMonoid M] [Module R M] --- TODO: show `[Algebra R A] : Algebra (Ideal R) A` too instance moduleSubmodule : Module (Ideal R) (Submodule R M) where smul_add := smul_sup add_smul := sup_smul @@ -1260,6 +1259,39 @@ theorem set_smul_top_eq_span (s : Set R) : s • ⊤ = Ideal.span s := (span_smul_eq s ⊤).symm.trans (Ideal.span s).mul_top +variable {A B} [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] + +open Submodule + +instance algebraIdeal : Algebra (Ideal R) (Submodule R A) where + __ := moduleSubmodule + toFun := map (Algebra.linearMap R A) + map_one' := by + rw [one_eq_span, map_span, Set.image_singleton, Algebra.linearMap_apply, map_one, one_eq_span] + map_mul' := (Submodule.map_mul · · <| Algebra.ofId R A) + map_zero' := map_bot _ + map_add' := (map_sup · · _) + commutes' I M := mul_comm_of_commute <| by rintro _ ⟨r, _, rfl⟩ a _; apply Algebra.commutes + smul_def' I M := le_antisymm (smul_le.mpr fun r hr a ha ↦ by + rw [Algebra.smul_def]; exact Submodule.mul_mem_mul ⟨r, hr, rfl⟩ ha) (Submodule.mul_le.mpr <| by + rintro _ ⟨r, hr, rfl⟩ a ha; rw [Algebra.linearMap_apply, ← Algebra.smul_def] + exact Submodule.smul_mem_smul hr ha) + +/-- `Submonoid.map` as an `AlgHom`, when applied to an `AlgHom`. -/ +@[simps!] def mapAlgHom (f : A →ₐ[R] B) : Submodule R A →ₐ[Ideal R] Submodule R B where + __ := mapHom f + commutes' I := (map_comp _ _ I).symm.trans (congr_arg (map · I) <| LinearMap.ext f.commutes) + +/-- `Submonoid.map` as an `AlgEquiv`, when applied to an `AlgEquiv`. -/ +-- TODO: when A, B noncommutative, still has `MulEquiv`. +@[simps!] def mapAlgEquiv (f : A ≃ₐ[R] B) : Submodule R A ≃ₐ[Ideal R] Submodule R B where + __ := mapAlgHom f + invFun := mapAlgHom f.symm + left_inv I := (map_comp _ _ I).symm.trans <| + (congr_arg (map · I) <| LinearMap.ext (f.left_inv ·)).trans (map_id I) + right_inv I := (map_comp _ _ I).symm.trans <| + (congr_arg (map · I) <| LinearMap.ext (f.right_inv ·)).trans (map_id I) + end Submodule instance {R} [Semiring R] : NonUnitalSubsemiringClass (Ideal R) R where From 62c53187b929efb5b80b94f9feff0979c34642ef Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Mon, 18 Nov 2024 16:07:35 +0000 Subject: [PATCH 078/829] fix: `linear_combination` with an additive constant (#19082) Fix a bug in one of the "off-label" use cases of `linear_combination`, which would not turn up in mathlib (because it throws a warning) but might confuse users in the wild. --- Mathlib/Tactic/LinearCombination.lean | 6 +++--- MathlibTest/linear_combination.lean | 4 ++++ 2 files changed, 7 insertions(+), 3 deletions(-) diff --git a/Mathlib/Tactic/LinearCombination.lean b/Mathlib/Tactic/LinearCombination.lean index f330e31c1553d..f1c65e78a805b 100644 --- a/Mathlib/Tactic/LinearCombination.lean +++ b/Mathlib/Tactic/LinearCombination.lean @@ -105,14 +105,14 @@ partial def expandLinearCombo (ty : Expr) (stx : Syntax.Term) : TermElabM Expand logWarningAt c "this constant has no effect on the linear combination; it can be dropped \ from the term" pure (.proof rel p) - | .const c, .proof rel p => + | .const c, .proof eq p => logWarningAt c "this constant has no effect on the linear combination; it can be dropped \ from the term" - .proof rel <$> ``(Eq.symm $p) + .proof eq <$> ``(Eq.symm $p) | .proof rel₁ p₁, .proof eq p₂ => let i := mkIdent <| Ineq.addRelRelData rel₁ eq .proof rel₁ <$> ``($i $p₁ (Eq.symm $p₂)) - | .proof _ _, .proof _ _ => + | _, .proof _ _ => throwError "coefficients of inequalities in 'linear_combination' must be nonnegative" | `(-$e) => do match ← expandLinearCombo ty e with diff --git a/MathlibTest/linear_combination.lean b/MathlibTest/linear_combination.lean index 70b1339f6e214..4aa730f384389 100644 --- a/MathlibTest/linear_combination.lean +++ b/MathlibTest/linear_combination.lean @@ -384,6 +384,10 @@ h2 : b ≥ 0 #guard_msgs in example (a b : ℚ) (h1 : a ≤ 1) (h2 : b ≥ 0) : a ≤ b := by linear_combination h1 +/-- error: coefficients of inequalities in 'linear_combination' must be nonnegative -/ +#guard_msgs in +example (x y : ℤ) (h : x ≤ y) : -x ≤ -y := by linear_combination 4 - h + /-! ### Nonlinear inequalities -/ example {a b : ℝ} (ha : 0 ≤ a) (hb : b < 1) : a * b ≤ a := by linear_combination a * hb From 277fb944c0d5c31c9fd4709eb16b1eca9c3f636c Mon Sep 17 00:00:00 2001 From: Pim Otte Date: Mon, 18 Nov 2024 16:07:37 +0000 Subject: [PATCH 079/829] feat(Combinatorics/SimpleGraph/Matching): add `IsPerfectMatching.toSubgraph_spanningCoe_iff` (#19094) Just a small lemma about perfect matchings and `spanningCoe`. In preparation for Tutte's theorem. --- Mathlib/Combinatorics/SimpleGraph/Matching.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Mathlib/Combinatorics/SimpleGraph/Matching.lean b/Mathlib/Combinatorics/SimpleGraph/Matching.lean index 5df846e190c17..14f2cbfc960a6 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Matching.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Matching.lean @@ -204,6 +204,11 @@ lemma IsPerfectMatching.induce_connectedComponent_isMatching (h : M.IsPerfectMat (c : ConnectedComponent G) : (M.induce c.supp).IsMatching := by simpa [h.2.verts_eq_univ] using h.1.induce_connectedComponent c +@[simp] +lemma IsPerfectMatching.toSubgraph_spanningCoe_iff (h : M.spanningCoe ≤ G') : + (G'.toSubgraph M.spanningCoe h).IsPerfectMatching ↔ M.IsPerfectMatching := by + simp only [isPerfectMatching_iff, toSubgraph_adj, spanningCoe_adj] + end Subgraph namespace ConnectedComponent From 9c84692f2bf8b3468d258abe9caf4654396e9941 Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Mon, 18 Nov 2024 16:51:24 +0000 Subject: [PATCH 080/829] chore: Split Mathlib.RingTheory.Ideal.Norm (#19211) We split `Mathlib.RingTheory.Ideal.Norm` into `Basic` and `RelNorm`. This is in preparation to generalize `Ideal.RelNorm` to nonfree extensions. --- Mathlib.lean | 3 +- Mathlib/NumberTheory/Cyclotomic/Rat.lean | 2 +- Mathlib/RingTheory/FractionalIdeal/Norm.lean | 3 +- .../Ideal/{Norm.lean => Norm/AbsNorm.lean} | 218 +--------------- Mathlib/RingTheory/Ideal/Norm/RelNorm.lean | 232 ++++++++++++++++++ 5 files changed, 241 insertions(+), 217 deletions(-) rename Mathlib/RingTheory/Ideal/{Norm.lean => Norm/AbsNorm.lean} (69%) create mode 100644 Mathlib/RingTheory/Ideal/Norm/RelNorm.lean diff --git a/Mathlib.lean b/Mathlib.lean index c6aea2fa69b6f..6f62181b1e265 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4251,7 +4251,8 @@ import Mathlib.RingTheory.Ideal.Maps import Mathlib.RingTheory.Ideal.Maximal import Mathlib.RingTheory.Ideal.MinimalPrime import Mathlib.RingTheory.Ideal.Nonunits -import Mathlib.RingTheory.Ideal.Norm +import Mathlib.RingTheory.Ideal.Norm.AbsNorm +import Mathlib.RingTheory.Ideal.Norm.RelNorm import Mathlib.RingTheory.Ideal.Operations import Mathlib.RingTheory.Ideal.Over import Mathlib.RingTheory.Ideal.Pointwise diff --git a/Mathlib/NumberTheory/Cyclotomic/Rat.lean b/Mathlib/NumberTheory/Cyclotomic/Rat.lean index eae82d89db9ff..2c898d2ea3c30 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Rat.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Rat.lean @@ -5,7 +5,7 @@ Authors: Riccardo Brasca -/ import Mathlib.NumberTheory.Cyclotomic.Discriminant import Mathlib.RingTheory.Polynomial.Eisenstein.IsIntegral -import Mathlib.RingTheory.Ideal.Norm +import Mathlib.RingTheory.Ideal.Norm.AbsNorm /-! # Ring of integers of `p ^ n`-th cyclotomic fields diff --git a/Mathlib/RingTheory/FractionalIdeal/Norm.lean b/Mathlib/RingTheory/FractionalIdeal/Norm.lean index b21d23be5b447..0285775c548a1 100644 --- a/Mathlib/RingTheory/FractionalIdeal/Norm.lean +++ b/Mathlib/RingTheory/FractionalIdeal/Norm.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Xavier Roblot -/ import Mathlib.RingTheory.FractionalIdeal.Basic -import Mathlib.RingTheory.Ideal.Norm +import Mathlib.RingTheory.Ideal.Norm.AbsNorm +import Mathlib.RingTheory.Localization.NormTrace /-! diff --git a/Mathlib/RingTheory/Ideal/Norm.lean b/Mathlib/RingTheory/Ideal/Norm/AbsNorm.lean similarity index 69% rename from Mathlib/RingTheory/Ideal/Norm.lean rename to Mathlib/RingTheory/Ideal/Norm/AbsNorm.lean index ae759f8605dbe..69f9fc150ca4f 100644 --- a/Mathlib/RingTheory/Ideal/Norm.lean +++ b/Mathlib/RingTheory/Ideal/Norm/AbsNorm.lean @@ -8,17 +8,16 @@ import Mathlib.Data.Int.AbsoluteValue import Mathlib.Data.Int.Associated import Mathlib.LinearAlgebra.FreeModule.Determinant import Mathlib.LinearAlgebra.FreeModule.IdealQuotient -import Mathlib.RingTheory.DedekindDomain.PID -import Mathlib.RingTheory.Localization.NormTrace +import Mathlib.RingTheory.DedekindDomain.Dvr +import Mathlib.RingTheory.DedekindDomain.Ideal +import Mathlib.RingTheory.Norm.Basic /-! # Ideal norms This file defines the absolute ideal norm `Ideal.absNorm (I : Ideal R) : ℕ` as the cardinality of -the quotient `R ⧸ I` (setting it to 0 if the cardinality is infinite), -and the relative ideal norm `Ideal.spanNorm R (I : Ideal S) : Ideal S` as the ideal spanned by -the norms of elements in `I`. +the quotient `R ⧸ I` (setting it to 0 if the cardinality is infinite). ## Main definitions @@ -26,10 +25,6 @@ the norms of elements in `I`. This maps `⊥` to `0` and `⊤` to `1`. * `Ideal.absNorm (I : Ideal R)`: the absolute ideal norm, defined as the cardinality of the quotient `R ⧸ I`, as a bundled monoid-with-zero homomorphism. - * `Ideal.spanNorm R (I : Ideal S)`: the ideal spanned by the norms of elements in `I`. - This is used to define `Ideal.relNorm`. - * `Ideal.relNorm R (I : Ideal S)`: the relative ideal norm as a bundled monoid-with-zero morphism, - defined as the ideal spanned by the norms of elements in `I`. ## Main results @@ -39,10 +34,8 @@ the norms of elements in `I`. of the basis change matrix * `Ideal.absNorm_span_singleton`: the ideal norm of a principal ideal is the norm of its generator - * `map_mul Ideal.relNorm`: multiplicativity of the relative ideal norm -/ - open scoped nonZeroDivisors section abs_norm @@ -460,206 +453,3 @@ end Ideal end RingOfIntegers end abs_norm - -section SpanNorm - -namespace Ideal - -open Submodule - -variable (R : Type*) [CommRing R] {S : Type*} [CommRing S] [Algebra R S] - -/-- `Ideal.spanNorm R (I : Ideal S)` is the ideal generated by mapping `Algebra.norm R` over `I`. - -See also `Ideal.relNorm`. --/ -def spanNorm (I : Ideal S) : Ideal R := - Ideal.span (Algebra.norm R '' (I : Set S)) - -@[simp] -theorem spanNorm_bot [Nontrivial S] [Module.Free R S] [Module.Finite R S] : - spanNorm R (⊥ : Ideal S) = ⊥ := span_eq_bot.mpr fun x hx => by simpa using hx - -variable {R} - -@[simp] -theorem spanNorm_eq_bot_iff [IsDomain R] [IsDomain S] [Module.Free R S] [Module.Finite R S] - {I : Ideal S} : spanNorm R I = ⊥ ↔ I = ⊥ := by - simp only [spanNorm, Ideal.span_eq_bot, Set.mem_image, SetLike.mem_coe, forall_exists_index, - and_imp, forall_apply_eq_imp_iff₂, - Algebra.norm_eq_zero_iff_of_basis (Module.Free.chooseBasis R S), @eq_bot_iff _ _ _ I, - SetLike.le_def] - rfl - -variable (R) - -theorem norm_mem_spanNorm {I : Ideal S} (x : S) (hx : x ∈ I) : Algebra.norm R x ∈ I.spanNorm R := - subset_span (Set.mem_image_of_mem _ hx) - -@[simp] -theorem spanNorm_singleton {r : S} : spanNorm R (span ({r} : Set S)) = span {Algebra.norm R r} := - le_antisymm - (span_le.mpr fun x hx => - mem_span_singleton.mpr - (by - obtain ⟨x, hx', rfl⟩ := (Set.mem_image _ _ _).mp hx - exact map_dvd _ (mem_span_singleton.mp hx'))) - ((span_singleton_le_iff_mem _).mpr (norm_mem_spanNorm _ _ (mem_span_singleton_self _))) - -@[simp] -theorem spanNorm_top : spanNorm R (⊤ : Ideal S) = ⊤ := by - -- Porting note: was - -- simp [← Ideal.span_singleton_one] - rw [← Ideal.span_singleton_one, spanNorm_singleton] - simp - -theorem map_spanNorm (I : Ideal S) {T : Type*} [CommRing T] (f : R →+* T) : - map f (spanNorm R I) = span (f ∘ Algebra.norm R '' (I : Set S)) := by - rw [spanNorm, map_span, Set.image_image] - -- Porting note: `Function.comp` reducibility - rfl - -@[mono] -theorem spanNorm_mono {I J : Ideal S} (h : I ≤ J) : spanNorm R I ≤ spanNorm R J := - Ideal.span_mono (Set.monotone_image h) - -theorem spanNorm_localization (I : Ideal S) [Module.Finite R S] [Module.Free R S] (M : Submonoid R) - {Rₘ : Type*} (Sₘ : Type*) [CommRing Rₘ] [Algebra R Rₘ] [CommRing Sₘ] [Algebra S Sₘ] - [Algebra Rₘ Sₘ] [Algebra R Sₘ] [IsScalarTower R Rₘ Sₘ] [IsScalarTower R S Sₘ] - [IsLocalization M Rₘ] [IsLocalization (Algebra.algebraMapSubmonoid S M) Sₘ] : - spanNorm Rₘ (I.map (algebraMap S Sₘ)) = (spanNorm R I).map (algebraMap R Rₘ) := by - cases subsingleton_or_nontrivial R - · haveI := IsLocalization.unique R Rₘ M - simp [eq_iff_true_of_subsingleton] - let b := Module.Free.chooseBasis R S - rw [map_spanNorm] - refine span_eq_span (Set.image_subset_iff.mpr ?_) (Set.image_subset_iff.mpr ?_) - · rintro a' ha' - simp only [Set.mem_preimage, submodule_span_eq, ← map_spanNorm, SetLike.mem_coe, - IsLocalization.mem_map_algebraMap_iff (Algebra.algebraMapSubmonoid S M) Sₘ, - IsLocalization.mem_map_algebraMap_iff M Rₘ, Prod.exists] at ha' ⊢ - obtain ⟨⟨a, ha⟩, ⟨_, ⟨s, hs, rfl⟩⟩, has⟩ := ha' - refine ⟨⟨Algebra.norm R a, norm_mem_spanNorm _ _ ha⟩, - ⟨s ^ Fintype.card (Module.Free.ChooseBasisIndex R S), pow_mem hs _⟩, ?_⟩ - simp only [Submodule.coe_mk, Subtype.coe_mk, map_pow] at has ⊢ - apply_fun Algebra.norm Rₘ at has - rwa [_root_.map_mul, ← IsScalarTower.algebraMap_apply, IsScalarTower.algebraMap_apply R Rₘ, - Algebra.norm_algebraMap_of_basis (b.localizationLocalization Rₘ M Sₘ), - Algebra.norm_localization R M a] at has - · intro a ha - rw [Set.mem_preimage, Function.comp_apply, ← Algebra.norm_localization (Sₘ := Sₘ) R M a] - exact subset_span (Set.mem_image_of_mem _ (mem_map_of_mem _ ha)) - -theorem spanNorm_mul_spanNorm_le (I J : Ideal S) : - spanNorm R I * spanNorm R J ≤ spanNorm R (I * J) := by - rw [spanNorm, spanNorm, spanNorm, Ideal.span_mul_span', ← Set.image_mul] - refine Ideal.span_mono (Set.monotone_image ?_) - rintro _ ⟨x, hxI, y, hyJ, rfl⟩ - exact Ideal.mul_mem_mul hxI hyJ - -/-- This condition `eq_bot_or_top` is equivalent to being a field. -However, `Ideal.spanNorm_mul_of_field` is harder to apply since we'd need to upgrade a `CommRing R` -instance to a `Field R` instance. -/ -theorem spanNorm_mul_of_bot_or_top [IsDomain R] [IsDomain S] [Module.Free R S] [Module.Finite R S] - (eq_bot_or_top : ∀ I : Ideal R, I = ⊥ ∨ I = ⊤) (I J : Ideal S) : - spanNorm R (I * J) = spanNorm R I * spanNorm R J := by - refine le_antisymm ?_ (spanNorm_mul_spanNorm_le R _ _) - cases' eq_bot_or_top (spanNorm R I) with hI hI - · rw [hI, spanNorm_eq_bot_iff.mp hI, bot_mul, spanNorm_bot] - exact bot_le - rw [hI, Ideal.top_mul] - cases' eq_bot_or_top (spanNorm R J) with hJ hJ - · rw [hJ, spanNorm_eq_bot_iff.mp hJ, mul_bot, spanNorm_bot] - rw [hJ] - exact le_top - -@[simp] -theorem spanNorm_mul_of_field {K : Type*} [Field K] [Algebra K S] [IsDomain S] [Module.Finite K S] - (I J : Ideal S) : spanNorm K (I * J) = spanNorm K I * spanNorm K J := - spanNorm_mul_of_bot_or_top K eq_bot_or_top I J - -variable [IsDomain R] [IsDomain S] [IsDedekindDomain R] [IsDedekindDomain S] -variable [Module.Finite R S] [Module.Free R S] - -/-- Multiplicativity of `Ideal.spanNorm`. simp-normal form is `map_mul (Ideal.relNorm R)`. -/ -theorem spanNorm_mul (I J : Ideal S) : spanNorm R (I * J) = spanNorm R I * spanNorm R J := by - nontriviality R - cases subsingleton_or_nontrivial S - · have : ∀ I : Ideal S, I = ⊤ := fun I => Subsingleton.elim I ⊤ - simp [this I, this J, this (I * J)] - refine eq_of_localization_maximal ?_ - intro P hP - by_cases hP0 : P = ⊥ - · subst hP0 - rw [spanNorm_mul_of_bot_or_top] - intro I - refine or_iff_not_imp_right.mpr fun hI => ?_ - exact (hP.eq_of_le hI bot_le).symm - let P' := Algebra.algebraMapSubmonoid S P.primeCompl - letI : Algebra (Localization.AtPrime P) (Localization P') := localizationAlgebra P.primeCompl S - haveI : IsScalarTower R (Localization.AtPrime P) (Localization P') := - IsScalarTower.of_algebraMap_eq (fun x => - (IsLocalization.map_eq (T := P') (Q := Localization P') P.primeCompl.le_comap_map x).symm) - have h : P' ≤ S⁰ := - map_le_nonZeroDivisors_of_injective _ (NoZeroSMulDivisors.algebraMap_injective _ _) - P.primeCompl_le_nonZeroDivisors - haveI : IsDomain (Localization P') := IsLocalization.isDomain_localization h - haveI : IsDedekindDomain (Localization P') := IsLocalization.isDedekindDomain S h _ - letI := Classical.decEq (Ideal (Localization P')) - haveI : IsPrincipalIdealRing (Localization P') := - IsDedekindDomain.isPrincipalIdealRing_localization_over_prime S P hP0 - rw [Ideal.map_mul, ← spanNorm_localization R I P.primeCompl (Localization P'), - ← spanNorm_localization R J P.primeCompl (Localization P'), - ← spanNorm_localization R (I * J) P.primeCompl (Localization P'), Ideal.map_mul, - ← (I.map _).span_singleton_generator, ← (J.map _).span_singleton_generator, - span_singleton_mul_span_singleton, spanNorm_singleton, spanNorm_singleton, - spanNorm_singleton, span_singleton_mul_span_singleton, _root_.map_mul] - -/-- The relative norm `Ideal.relNorm R (I : Ideal S)`, where `R` and `S` are Dedekind domains, -and `S` is an extension of `R` that is finite and free as a module. -/ -def relNorm : Ideal S →*₀ Ideal R where - toFun := spanNorm R - map_zero' := spanNorm_bot R - map_one' := by dsimp only; rw [one_eq_top, spanNorm_top R, one_eq_top] - map_mul' := spanNorm_mul R - -theorem relNorm_apply (I : Ideal S) : relNorm R I = span (Algebra.norm R '' (I : Set S) : Set R) := - rfl - -@[simp] -theorem spanNorm_eq (I : Ideal S) : spanNorm R I = relNorm R I := rfl - -@[simp] -theorem relNorm_bot : relNorm R (⊥ : Ideal S) = ⊥ := by - simpa only [zero_eq_bot] using map_zero (relNorm R : Ideal S →*₀ _) - -@[simp] -theorem relNorm_top : relNorm R (⊤ : Ideal S) = ⊤ := by - simpa only [one_eq_top] using map_one (relNorm R : Ideal S →*₀ _) - -variable {R} - -@[simp] -theorem relNorm_eq_bot_iff {I : Ideal S} : relNorm R I = ⊥ ↔ I = ⊥ := - spanNorm_eq_bot_iff - -variable (R) - -theorem norm_mem_relNorm (I : Ideal S) {x : S} (hx : x ∈ I) : Algebra.norm R x ∈ relNorm R I := - norm_mem_spanNorm R x hx - -@[simp] -theorem relNorm_singleton (r : S) : relNorm R (span ({r} : Set S)) = span {Algebra.norm R r} := - spanNorm_singleton R - -theorem map_relNorm (I : Ideal S) {T : Type*} [CommRing T] (f : R →+* T) : - map f (relNorm R I) = span (f ∘ Algebra.norm R '' (I : Set S)) := - map_spanNorm R I f - -@[mono] -theorem relNorm_mono {I J : Ideal S} (h : I ≤ J) : relNorm R I ≤ relNorm R J := - spanNorm_mono R h - -end Ideal - -end SpanNorm diff --git a/Mathlib/RingTheory/Ideal/Norm/RelNorm.lean b/Mathlib/RingTheory/Ideal/Norm/RelNorm.lean new file mode 100644 index 0000000000000..ccb7c0362da1e --- /dev/null +++ b/Mathlib/RingTheory/Ideal/Norm/RelNorm.lean @@ -0,0 +1,232 @@ +/- +Copyright (c) 2022 Anne Baanen. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anne Baanen, Alex J. Best +-/ +import Mathlib.LinearAlgebra.FreeModule.PID +import Mathlib.RingTheory.DedekindDomain.PID +import Mathlib.RingTheory.Localization.NormTrace + +/-! + +# Ideal norms + +This file defines the relative ideal norm `Ideal.spanNorm R (I : Ideal S) : Ideal S` as the ideal +spanned by the norms of elements in `I`. + +## Main definitions + + * `Ideal.spanNorm R (I : Ideal S)`: the ideal spanned by the norms of elements in `I`. + This is used to define `Ideal.relNorm`. + * `Ideal.relNorm R (I : Ideal S)`: the relative ideal norm as a bundled monoid-with-zero morphism, + defined as the ideal spanned by the norms of elements in `I`. + +## Main results + + * `map_mul Ideal.relNorm`: multiplicativity of the relative ideal norm +-/ + +open scoped nonZeroDivisors + +section SpanNorm + +namespace Ideal + +open Submodule + +variable (R : Type*) [CommRing R] {S : Type*} [CommRing S] [Algebra R S] + +/-- `Ideal.spanNorm R (I : Ideal S)` is the ideal generated by mapping `Algebra.norm R` over `I`. + +See also `Ideal.relNorm`. +-/ +def spanNorm (I : Ideal S) : Ideal R := + Ideal.span (Algebra.norm R '' (I : Set S)) + +@[simp] +theorem spanNorm_bot [Nontrivial S] [Module.Free R S] [Module.Finite R S] : + spanNorm R (⊥ : Ideal S) = ⊥ := span_eq_bot.mpr fun x hx => by simpa using hx + +variable {R} + +@[simp] +theorem spanNorm_eq_bot_iff [IsDomain R] [IsDomain S] [Module.Free R S] [Module.Finite R S] + {I : Ideal S} : spanNorm R I = ⊥ ↔ I = ⊥ := by + simp only [spanNorm, Ideal.span_eq_bot, Set.mem_image, SetLike.mem_coe, forall_exists_index, + and_imp, forall_apply_eq_imp_iff₂, + Algebra.norm_eq_zero_iff_of_basis (Module.Free.chooseBasis R S), @eq_bot_iff _ _ _ I, + SetLike.le_def] + rfl + +variable (R) + +theorem norm_mem_spanNorm {I : Ideal S} (x : S) (hx : x ∈ I) : Algebra.norm R x ∈ I.spanNorm R := + subset_span (Set.mem_image_of_mem _ hx) + +@[simp] +theorem spanNorm_singleton {r : S} : spanNorm R (span ({r} : Set S)) = span {Algebra.norm R r} := + le_antisymm + (span_le.mpr fun x hx => + mem_span_singleton.mpr + (by + obtain ⟨x, hx', rfl⟩ := (Set.mem_image _ _ _).mp hx + exact map_dvd _ (mem_span_singleton.mp hx'))) + ((span_singleton_le_iff_mem _).mpr (norm_mem_spanNorm _ _ (mem_span_singleton_self _))) + +@[simp] +theorem spanNorm_top : spanNorm R (⊤ : Ideal S) = ⊤ := by + -- Porting note: was + -- simp [← Ideal.span_singleton_one] + rw [← Ideal.span_singleton_one, spanNorm_singleton] + simp + +theorem map_spanNorm (I : Ideal S) {T : Type*} [CommRing T] (f : R →+* T) : + map f (spanNorm R I) = span (f ∘ Algebra.norm R '' (I : Set S)) := by + rw [spanNorm, map_span, Set.image_image] + -- Porting note: `Function.comp` reducibility + rfl + +@[mono] +theorem spanNorm_mono {I J : Ideal S} (h : I ≤ J) : spanNorm R I ≤ spanNorm R J := + Ideal.span_mono (Set.monotone_image h) + +theorem spanNorm_localization (I : Ideal S) [Module.Finite R S] [Module.Free R S] (M : Submonoid R) + {Rₘ : Type*} (Sₘ : Type*) [CommRing Rₘ] [Algebra R Rₘ] [CommRing Sₘ] [Algebra S Sₘ] + [Algebra Rₘ Sₘ] [Algebra R Sₘ] [IsScalarTower R Rₘ Sₘ] [IsScalarTower R S Sₘ] + [IsLocalization M Rₘ] [IsLocalization (Algebra.algebraMapSubmonoid S M) Sₘ] : + spanNorm Rₘ (I.map (algebraMap S Sₘ)) = (spanNorm R I).map (algebraMap R Rₘ) := by + cases subsingleton_or_nontrivial R + · haveI := IsLocalization.unique R Rₘ M + simp [eq_iff_true_of_subsingleton] + let b := Module.Free.chooseBasis R S + rw [map_spanNorm] + refine span_eq_span (Set.image_subset_iff.mpr ?_) (Set.image_subset_iff.mpr ?_) + · rintro a' ha' + simp only [Set.mem_preimage, submodule_span_eq, ← map_spanNorm, SetLike.mem_coe, + IsLocalization.mem_map_algebraMap_iff (Algebra.algebraMapSubmonoid S M) Sₘ, + IsLocalization.mem_map_algebraMap_iff M Rₘ, Prod.exists] at ha' ⊢ + obtain ⟨⟨a, ha⟩, ⟨_, ⟨s, hs, rfl⟩⟩, has⟩ := ha' + refine ⟨⟨Algebra.norm R a, norm_mem_spanNorm _ _ ha⟩, + ⟨s ^ Fintype.card (Module.Free.ChooseBasisIndex R S), pow_mem hs _⟩, ?_⟩ + simp only [Submodule.coe_mk, Subtype.coe_mk, map_pow] at has ⊢ + apply_fun Algebra.norm Rₘ at has + rwa [_root_.map_mul, ← IsScalarTower.algebraMap_apply, IsScalarTower.algebraMap_apply R Rₘ, + Algebra.norm_algebraMap_of_basis (b.localizationLocalization Rₘ M Sₘ), + Algebra.norm_localization R M a] at has + · intro a ha + rw [Set.mem_preimage, Function.comp_apply, ← Algebra.norm_localization (Sₘ := Sₘ) R M a] + exact subset_span (Set.mem_image_of_mem _ (mem_map_of_mem _ ha)) + +theorem spanNorm_mul_spanNorm_le (I J : Ideal S) : + spanNorm R I * spanNorm R J ≤ spanNorm R (I * J) := by + rw [spanNorm, spanNorm, spanNorm, Ideal.span_mul_span', ← Set.image_mul] + refine Ideal.span_mono (Set.monotone_image ?_) + rintro _ ⟨x, hxI, y, hyJ, rfl⟩ + exact Ideal.mul_mem_mul hxI hyJ + +/-- This condition `eq_bot_or_top` is equivalent to being a field. +However, `Ideal.spanNorm_mul_of_field` is harder to apply since we'd need to upgrade a `CommRing R` +instance to a `Field R` instance. -/ +theorem spanNorm_mul_of_bot_or_top [IsDomain R] [IsDomain S] [Module.Free R S] [Module.Finite R S] + (eq_bot_or_top : ∀ I : Ideal R, I = ⊥ ∨ I = ⊤) (I J : Ideal S) : + spanNorm R (I * J) = spanNorm R I * spanNorm R J := by + refine le_antisymm ?_ (spanNorm_mul_spanNorm_le R _ _) + cases' eq_bot_or_top (spanNorm R I) with hI hI + · rw [hI, spanNorm_eq_bot_iff.mp hI, bot_mul, spanNorm_bot] + exact bot_le + rw [hI, Ideal.top_mul] + cases' eq_bot_or_top (spanNorm R J) with hJ hJ + · rw [hJ, spanNorm_eq_bot_iff.mp hJ, mul_bot, spanNorm_bot] + rw [hJ] + exact le_top + +@[simp] +theorem spanNorm_mul_of_field {K : Type*} [Field K] [Algebra K S] [IsDomain S] [Module.Finite K S] + (I J : Ideal S) : spanNorm K (I * J) = spanNorm K I * spanNorm K J := + spanNorm_mul_of_bot_or_top K eq_bot_or_top I J + +variable [IsDomain R] [IsDomain S] [IsDedekindDomain R] [IsDedekindDomain S] +variable [Module.Finite R S] [Module.Free R S] + +/-- Multiplicativity of `Ideal.spanNorm`. simp-normal form is `map_mul (Ideal.relNorm R)`. -/ +theorem spanNorm_mul (I J : Ideal S) : spanNorm R (I * J) = spanNorm R I * spanNorm R J := by + nontriviality R + cases subsingleton_or_nontrivial S + · have : ∀ I : Ideal S, I = ⊤ := fun I => Subsingleton.elim I ⊤ + simp [this I, this J, this (I * J)] + refine eq_of_localization_maximal ?_ + intro P hP + by_cases hP0 : P = ⊥ + · subst hP0 + rw [spanNorm_mul_of_bot_or_top] + intro I + refine or_iff_not_imp_right.mpr fun hI => ?_ + exact (hP.eq_of_le hI bot_le).symm + let P' := Algebra.algebraMapSubmonoid S P.primeCompl + letI : Algebra (Localization.AtPrime P) (Localization P') := localizationAlgebra P.primeCompl S + haveI : IsScalarTower R (Localization.AtPrime P) (Localization P') := + IsScalarTower.of_algebraMap_eq (fun x => + (IsLocalization.map_eq (T := P') (Q := Localization P') P.primeCompl.le_comap_map x).symm) + have h : P' ≤ S⁰ := + map_le_nonZeroDivisors_of_injective _ (NoZeroSMulDivisors.algebraMap_injective _ _) + P.primeCompl_le_nonZeroDivisors + haveI : IsDomain (Localization P') := IsLocalization.isDomain_localization h + haveI : IsDedekindDomain (Localization P') := IsLocalization.isDedekindDomain S h _ + letI := Classical.decEq (Ideal (Localization P')) + haveI : IsPrincipalIdealRing (Localization P') := + IsDedekindDomain.isPrincipalIdealRing_localization_over_prime S P hP0 + rw [Ideal.map_mul, ← spanNorm_localization R I P.primeCompl (Localization P'), + ← spanNorm_localization R J P.primeCompl (Localization P'), + ← spanNorm_localization R (I * J) P.primeCompl (Localization P'), Ideal.map_mul, + ← (I.map _).span_singleton_generator, ← (J.map _).span_singleton_generator, + span_singleton_mul_span_singleton, spanNorm_singleton, spanNorm_singleton, + spanNorm_singleton, span_singleton_mul_span_singleton, _root_.map_mul] + +/-- The relative norm `Ideal.relNorm R (I : Ideal S)`, where `R` and `S` are Dedekind domains, +and `S` is an extension of `R` that is finite and free as a module. -/ +def relNorm : Ideal S →*₀ Ideal R where + toFun := spanNorm R + map_zero' := spanNorm_bot R + map_one' := by dsimp only; rw [one_eq_top, spanNorm_top R, one_eq_top] + map_mul' := spanNorm_mul R + +theorem relNorm_apply (I : Ideal S) : relNorm R I = span (Algebra.norm R '' (I : Set S) : Set R) := + rfl + +@[simp] +theorem spanNorm_eq (I : Ideal S) : spanNorm R I = relNorm R I := rfl + +@[simp] +theorem relNorm_bot : relNorm R (⊥ : Ideal S) = ⊥ := by + simpa only [zero_eq_bot] using map_zero (relNorm R : Ideal S →*₀ _) + +@[simp] +theorem relNorm_top : relNorm R (⊤ : Ideal S) = ⊤ := by + simpa only [one_eq_top] using map_one (relNorm R : Ideal S →*₀ _) + +variable {R} + +@[simp] +theorem relNorm_eq_bot_iff {I : Ideal S} : relNorm R I = ⊥ ↔ I = ⊥ := + spanNorm_eq_bot_iff + +variable (R) + +theorem norm_mem_relNorm (I : Ideal S) {x : S} (hx : x ∈ I) : Algebra.norm R x ∈ relNorm R I := + norm_mem_spanNorm R x hx + +@[simp] +theorem relNorm_singleton (r : S) : relNorm R (span ({r} : Set S)) = span {Algebra.norm R r} := + spanNorm_singleton R + +theorem map_relNorm (I : Ideal S) {T : Type*} [CommRing T] (f : R →+* T) : + map f (relNorm R I) = span (f ∘ Algebra.norm R '' (I : Set S)) := + map_spanNorm R I f + +@[mono] +theorem relNorm_mono {I J : Ideal S} (h : I ≤ J) : relNorm R I ≤ relNorm R J := + spanNorm_mono R h + +end Ideal + +end SpanNorm From a77f59fa3850f59f2eece0a4b1a332e75f25e7a0 Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 18 Nov 2024 17:00:56 +0000 Subject: [PATCH 081/829] fix: rename DeprecateMe to DeprecateTo and remove import (#19151) [Zulip report](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/DeprecateMe/near/482849897) --- Mathlib.lean | 2 +- Mathlib/Tactic.lean | 2 +- Mathlib/Tactic/Common.lean | 2 +- Mathlib/Tactic/{DeprecateMe.lean => DeprecateTo.lean} | 4 ++-- MathlibTest/{DeprecateMe.lean => DeprecateTo.lean} | 3 +-- 5 files changed, 6 insertions(+), 7 deletions(-) rename Mathlib/Tactic/{DeprecateMe.lean => DeprecateTo.lean} (98%) rename MathlibTest/{DeprecateMe.lean => DeprecateTo.lean} (91%) diff --git a/Mathlib.lean b/Mathlib.lean index 6f62181b1e265..38f129ac3f80a 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4611,7 +4611,7 @@ import Mathlib.Tactic.Convert import Mathlib.Tactic.Core import Mathlib.Tactic.DeclarationNames import Mathlib.Tactic.DefEqTransformations -import Mathlib.Tactic.DeprecateMe +import Mathlib.Tactic.DeprecateTo import Mathlib.Tactic.DeriveFintype import Mathlib.Tactic.DeriveToExpr import Mathlib.Tactic.DeriveTraversable diff --git a/Mathlib/Tactic.lean b/Mathlib/Tactic.lean index cbf02e204731c..a51290ea7ff52 100644 --- a/Mathlib/Tactic.lean +++ b/Mathlib/Tactic.lean @@ -64,7 +64,7 @@ import Mathlib.Tactic.Convert import Mathlib.Tactic.Core import Mathlib.Tactic.DeclarationNames import Mathlib.Tactic.DefEqTransformations -import Mathlib.Tactic.DeprecateMe +import Mathlib.Tactic.DeprecateTo import Mathlib.Tactic.DeriveFintype import Mathlib.Tactic.DeriveToExpr import Mathlib.Tactic.DeriveTraversable diff --git a/Mathlib/Tactic/Common.lean b/Mathlib/Tactic/Common.lean index 3ce8555b57c18..90340cd10a34e 100644 --- a/Mathlib/Tactic/Common.lean +++ b/Mathlib/Tactic/Common.lean @@ -45,7 +45,7 @@ import Mathlib.Tactic.Contrapose import Mathlib.Tactic.Conv import Mathlib.Tactic.Convert import Mathlib.Tactic.DefEqTransformations -import Mathlib.Tactic.DeprecateMe +import Mathlib.Tactic.DeprecateTo import Mathlib.Tactic.DeriveToExpr import Mathlib.Tactic.Eqns import Mathlib.Tactic.ExistsI diff --git a/Mathlib/Tactic/DeprecateMe.lean b/Mathlib/Tactic/DeprecateTo.lean similarity index 98% rename from Mathlib/Tactic/DeprecateMe.lean rename to Mathlib/Tactic/DeprecateTo.lean index b84b97f0b4c2c..4f8d8522506e0 100644 --- a/Mathlib/Tactic/DeprecateMe.lean +++ b/Mathlib/Tactic/DeprecateTo.lean @@ -36,7 +36,7 @@ TODO: * preserve formatting of existing command? -/ -namespace Mathlib.Tactic.DeprecateMe +namespace Mathlib.Tactic.DeprecateTo open Lean Elab Term Command @@ -147,4 +147,4 @@ elab tk:"deprecate to " id:ident* dat:(ppSpace str ppSpace)? ppLine cmd:command addSuggestion (header := msg ++ "\n\nTry this:\n") (← getRef) toMessageData -end Mathlib.Tactic.DeprecateMe +end Mathlib.Tactic.DeprecateTo diff --git a/MathlibTest/DeprecateMe.lean b/MathlibTest/DeprecateTo.lean similarity index 91% rename from MathlibTest/DeprecateMe.lean rename to MathlibTest/DeprecateTo.lean index d745326c337fa..4c4263b6c0773 100644 --- a/MathlibTest/DeprecateMe.lean +++ b/MathlibTest/DeprecateTo.lean @@ -1,6 +1,5 @@ -import Mathlib.Tactic.DeprecateMe +import Mathlib.Tactic.DeprecateTo import Mathlib.Tactic.ToAdditive -import Mathlib.Tactic /-- info: * Pairings: From 2ff73d5985b6084b9a7b3f402b90de35951c3e1c Mon Sep 17 00:00:00 2001 From: Yoh Tanimoto Date: Mon, 18 Nov 2024 18:16:53 +0000 Subject: [PATCH 082/829] =?UTF-8?q?feat(Mathlib/Algebra/Order):=20add=20eq?= =?UTF-8?q?uivalent=20conditions=20to=20`a=20=E2=89=A4=20b`=20(#19166)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Two theorems that give equivalent conditions to `a ≤ b` in terms of `ε`-approximation. - Generalize `a ≤ b ↔ ∀ ε, 1 < ε → a ≤ b * ε` (`le_iff_forall_one_lt_le_mul`) to monoids. - Add `(hb : 0 ≤ b) : a ≤ b ↔ ∀ ε, 1 < ε → a ≤ b * ε` for linearly ordered semifields. the proofs were suggested by @fpvandoorn in [this thread](https://leanprover.zulipchat.com/#narrow/channel/217875-Is-there-code-for-X.3F/topic/.60le_iff_forall_pos_le_add.60.20for.20.60NNReal.60) motivation: I will need them in the proof of regularity of `rieszContent`, that will be defined after #18775 Co-authored-by: Yoh Tanimoto <57562556+yoh-tanimoto@users.noreply.github.com> --- Mathlib/Algebra/Order/Field/Basic.lean | 10 ++++++++++ Mathlib/Algebra/Order/Group/DenselyOrdered.lean | 4 ---- Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean | 5 +++++ 3 files changed, 15 insertions(+), 4 deletions(-) diff --git a/Mathlib/Algebra/Order/Field/Basic.lean b/Mathlib/Algebra/Order/Field/Basic.lean index b7d6f496a867d..6d24b4b52ebbe 100644 --- a/Mathlib/Algebra/Order/Field/Basic.lean +++ b/Mathlib/Algebra/Order/Field/Basic.lean @@ -435,6 +435,16 @@ theorem inv_pow_anti (a1 : 1 ≤ a) : Antitone fun n : ℕ => (a ^ n)⁻¹ := fu theorem inv_pow_strictAnti (a1 : 1 < a) : StrictAnti fun n : ℕ => (a ^ n)⁻¹ := fun _ _ => inv_pow_lt_inv_pow_of_lt a1 +theorem le_iff_forall_one_lt_le_mul₀ {α : Type*} [LinearOrderedSemifield α] + {a b : α} (hb : 0 ≤ b) : a ≤ b ↔ ∀ ε, 1 < ε → a ≤ b * ε := by + refine ⟨fun h _ hε ↦ h.trans <| le_mul_of_one_le_right hb hε.le, fun h ↦ ?_⟩ + obtain rfl|hb := hb.eq_or_lt + · simp_rw [zero_mul] at h + exact h 2 one_lt_two + refine le_of_forall_le_of_dense fun x hbx => ?_ + convert h (x / b) ((one_lt_div hb).mpr hbx) + rw [mul_div_cancel₀ _ hb.ne'] + /-! ### Results about `IsGLB` -/ diff --git a/Mathlib/Algebra/Order/Group/DenselyOrdered.lean b/Mathlib/Algebra/Order/Group/DenselyOrdered.lean index 41c6ac74c506b..a019c1739800a 100644 --- a/Mathlib/Algebra/Order/Group/DenselyOrdered.lean +++ b/Mathlib/Algebra/Order/Group/DenselyOrdered.lean @@ -28,10 +28,6 @@ theorem le_of_forall_one_lt_div_le (h : ∀ ε : α, 1 < ε → a / ε ≤ b) : le_of_forall_lt_one_mul_le fun ε ε1 => by simpa only [div_eq_mul_inv, inv_inv] using h ε⁻¹ (Left.one_lt_inv_iff.2 ε1) -@[to_additive] -theorem le_iff_forall_one_lt_le_mul : a ≤ b ↔ ∀ ε, 1 < ε → a ≤ b * ε := - ⟨fun h _ ε_pos => le_mul_of_le_of_one_le h ε_pos.le, le_of_forall_one_lt_le_mul⟩ - @[to_additive] theorem le_iff_forall_lt_one_mul_le : a ≤ b ↔ ∀ ε < 1, a * ε ≤ b := le_iff_forall_one_lt_le_mul (α := αᵒᵈ) diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean index 4868b5b043691..1bd716addf327 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean @@ -82,4 +82,9 @@ theorem le_iff_forall_one_lt_lt_mul' [MulLeftStrictMono α] : a ≤ b ↔ ∀ ε, 1 < ε → a < b * ε := ⟨fun h _ => lt_mul_of_le_of_one_lt h, le_of_forall_one_lt_lt_mul'⟩ +@[to_additive] +theorem le_iff_forall_one_lt_le_mul [MulLeftStrictMono α] : + a ≤ b ↔ ∀ ε, 1 < ε → a ≤ b * ε := + ⟨fun h _ hε ↦ lt_mul_of_le_of_one_lt h hε |>.le, le_of_forall_one_lt_le_mul⟩ + end ExistsMulOfLE From 7c8674a2fdb2b927c142cc964cba74df0d1066b5 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Mon, 18 Nov 2024 18:54:15 +0000 Subject: [PATCH 083/829] feat: sufficient condition for a presheaf to preserve finite limits (#18283) Co-authored-by: Markus Himmel --- Mathlib.lean | 1 + .../Limits/Preserves/Presheaf.lean | 189 ++++++++++++++++++ 2 files changed, 190 insertions(+) create mode 100644 Mathlib/CategoryTheory/Limits/Preserves/Presheaf.lean diff --git a/Mathlib.lean b/Mathlib.lean index 38f129ac3f80a..2314aad484f69 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1795,6 +1795,7 @@ import Mathlib.CategoryTheory.Limits.Preserves.Finite import Mathlib.CategoryTheory.Limits.Preserves.FunctorCategory import Mathlib.CategoryTheory.Limits.Preserves.Limits import Mathlib.CategoryTheory.Limits.Preserves.Opposites +import Mathlib.CategoryTheory.Limits.Preserves.Presheaf import Mathlib.CategoryTheory.Limits.Preserves.Shapes.BinaryProducts import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Biproducts import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Equalizers diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Presheaf.lean b/Mathlib/CategoryTheory/Limits/Preserves/Presheaf.lean new file mode 100644 index 0000000000000..f2a3c01bd9fc0 --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Preserves/Presheaf.lean @@ -0,0 +1,189 @@ +/- +Copyright (c) 2024 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel +-/ +import Mathlib.CategoryTheory.Limits.FilteredColimitCommutesFiniteLimit +import Mathlib.CategoryTheory.Limits.Elements + +/-! +# Finite-limit-preserving presheaves + +In this file we prove that if `C` is a small finitely cocomplete category and `A : Cᵒᵖ ⥤ Type u` +is a presheaf, then `CostructuredArrow yoneda A` is filtered (equivalently, the category of +elements of `A` is cofiltered) if and only if `A` preserves finite limits. + +This is one of the keys steps of establishing the equivalence `Ind C ≌ (Cᵒᵖ ⥤ₗ Type u)` (here, +`Cᵒᵖ ⥤ₗ Type u` is the category of left exact functors) for a *small* finitely cocomplete category +`C`. + +## Implementation notes + +In the entire file, we are careful about details like functor associativity to ensure that we do +not end up in situations where we have morphisms like `colimit.ι F i ≫ f`, where the domain of `f` +is `colimit G` where `F` and `G` are definitionally equal but not syntactically equal. In these +situations, lemmas about `colimit.ι G i ≫ f` cannot be applied using `rw` and `simp`, forcing the +use of `erw`. In particular, for us this means that `HasColimit.isoOfNatIso (Iso.refl _)` is better +than `Iso.refl _` and that we should be very particular about functor associativity. + +The theorem is also true for large categories and the proof given here generalizes to this case on +paper. However, our infrastructure for commuting finite limits of shape `J` and filtered colimits +of shape `K` is fundamentally not equipped to deal with the case where not all limits of shape `K` +exist. In fact, not even the definition `colimitLimitToLimitColimit` can be written down if not +all limits of shape `K` exist. Refactoring this to the level of generality we need would be a major +undertaking. Since the application to the category of Ind-objects only require the case where `C` +is small, we leave this as a TODO. + +## References +* [M. Kashiwara, P. Schapira, *Categories and Sheaves*][Kashiwara2006], Proposition 3.3.13 +* [F. Borceux, *Handbook of Categorical Algebra 1*][borceux-vol1], Proposition 6.1.2 +-/ + +open CategoryTheory Limits + +universe v u + +namespace CategoryTheory.Limits + +section LargeCategory + +variable {C : Type u} [Category.{v} C] [HasFiniteColimits C] (A : Cᵒᵖ ⥤ Type v) + +/-- If `C` is a finitely cocomplete category and `A : Cᵒᵖ ⥤ Type u` is a presheaf that preserves +finite limites, then `CostructuredArrow yoneda A` is filtered. + +One direction of Proposition 3.3.13 of [Kashiwara2006]. +-/ +theorem isFiltered_costructuredArrow_yoneda_of_preservesFiniteLimits + [PreservesFiniteLimits A] : IsFiltered (CostructuredArrow yoneda A) := by + suffices IsCofiltered A.Elements from + IsFiltered.of_equivalence (CategoryOfElements.costructuredArrowYonedaEquivalence _) + suffices HasFiniteLimits A.Elements from IsCofiltered.of_hasFiniteLimits A.Elements + exact ⟨fun J _ _ => inferInstance⟩ + +end LargeCategory + +variable {C : Type u} [SmallCategory C] [HasFiniteColimits C] + +variable (A : Cᵒᵖ ⥤ Type u) + +namespace PreservesFiniteLimitsOfIsFilteredCostructuredArrowYonedaAux + +variable {J : Type} [SmallCategory J] [FinCategory J] (K : J ⥤ Cᵒᵖ) + +/-- (Implementation) This is the bifunctor we will apply "filtered colimits commute with finite +limits" to. -/ +def functorToInterchange : J ⥤ CostructuredArrow yoneda A ⥤ Type u := + K ⋙ coyoneda ⋙ (whiskeringLeft _ _ _).obj (CostructuredArrow.proj _ _) + +/-- (Implementation) The definition of `functorToInterchange`. -/ +@[simps!] +def functorToInterchangeIso : functorToInterchange A K ≅ + K ⋙ coyoneda ⋙ (whiskeringLeft _ _ _).obj (CostructuredArrow.proj _ _) := + Iso.refl _ + +/-- (Implementation) One way to express the flipped version of our functor. We choose this +association because the type of `Presheaf.tautologicalCocone` is +`Cocone (CostructuredArrow.proj yoneda P ⋙ yoneda)`, so this association will show up in the +proof.-/ +@[simps!] +def flipFunctorToInterchange : (functorToInterchange A K).flip ≅ + ((CostructuredArrow.proj yoneda A ⋙ yoneda) ⋙ (whiskeringLeft J Cᵒᵖ (Type u)).obj K) := + Iso.refl _ + +/-- (Implementation) A natural isomorphism we will need to construct `iso`. -/ +@[simps! (config := { fullyApplied := false }) hom_app] +noncomputable def isoAux : + (CostructuredArrow.proj yoneda A ⋙ yoneda ⋙ (evaluation Cᵒᵖ (Type u)).obj (limit K)) ≅ + ((coyoneda ⋙ (whiskeringLeft (CostructuredArrow yoneda A) C (Type u)).obj + (CostructuredArrow.proj yoneda A)).obj (limit K)) := + Iso.refl _ + +/-- (Implementation) The isomorphism that proves that `A` preserves finite limits. -/ +noncomputable def iso [IsFiltered (CostructuredArrow yoneda A)] : + A.obj (limit K) ≅ limit (K ⋙ A) := calc + A.obj (limit K) ≅ (colimit (CostructuredArrow.proj yoneda A ⋙ yoneda)).obj (limit K) := + (IsColimit.coconePointUniqueUpToIso + (Presheaf.isColimitTautologicalCocone A) (colimit.isColimit _)).app _ + _ ≅ colimit (((CostructuredArrow.proj yoneda A) ⋙ yoneda) ⋙ (evaluation _ _).obj (limit K)) := + (colimitObjIsoColimitCompEvaluation _ _) + _ ≅ colimit ((CostructuredArrow.proj yoneda A) ⋙ yoneda ⋙ (evaluation _ _).obj (limit K)) := + HasColimit.isoOfNatIso (Functor.associator _ _ _) + _ ≅ colimit + ((coyoneda ⋙ (whiskeringLeft _ _ _).obj (CostructuredArrow.proj yoneda A)).obj (limit K)) := + HasColimit.isoOfNatIso (isoAux A K) + _ ≅ colimit (limit (K ⋙ (coyoneda ⋙ (whiskeringLeft _ _ _).obj (CostructuredArrow.proj _ _)))) := + HasColimit.isoOfNatIso (preservesLimitIso _ _) + _ ≅ colimit (limit (functorToInterchange A K)) := + HasColimit.isoOfNatIso (HasLimit.isoOfNatIso (functorToInterchangeIso A K).symm) + _ ≅ limit (colimit (functorToInterchange A K).flip) := colimitLimitIso _ + _ ≅ limit (colimit + ((CostructuredArrow.proj yoneda A ⋙ yoneda) ⋙ (whiskeringLeft _ _ _).obj K)) := + HasLimit.isoOfNatIso (HasColimit.isoOfNatIso (flipFunctorToInterchange A K)) + _ ≅ limit (K ⋙ colimit (CostructuredArrow.proj yoneda A ⋙ yoneda)) := + HasLimit.isoOfNatIso + (colimitCompWhiskeringLeftIsoCompColimit (CostructuredArrow.proj yoneda A ⋙ yoneda) K) + _ ≅ limit (K ⋙ A) := HasLimit.isoOfNatIso (isoWhiskerLeft _ + (IsColimit.coconePointUniqueUpToIso + (colimit.isColimit _) (Presheaf.isColimitTautologicalCocone A))) + +theorem iso_hom [IsFiltered (CostructuredArrow yoneda A)] : (iso A K).hom = limit.post K A := by + -- We will have to use `ι_colimitLimitIso_limit_π` eventually, so let's start by + -- transforming the goal into something from a colimit to a limit so that we can apply + -- `limit.hom_ext` and `colimit.hom_ext`. + dsimp [iso, -Iso.app_hom] + simp only [Category.assoc] + rw [Eq.comm, ← Iso.inv_comp_eq, ← Iso.inv_comp_eq] + refine limit.hom_ext (fun j => colimit.hom_ext (fun i => ?_)) + simp only [Category.assoc] + + -- `simp` is not too helpful here because we will need to apply `NatTrans.comp_app_assoc` + -- backwards at certain points, so we rewrite the term manually. + rw [HasLimit.isoOfNatIso_hom_π, HasLimit.isoOfNatIso_hom_π_assoc, limit.post_π, + colimitObjIsoColimitCompEvaluation_ι_inv_assoc (CostructuredArrow.proj yoneda A ⋙ yoneda), + Iso.app_inv, ← NatTrans.comp_app_assoc, colimit.comp_coconePointUniqueUpToIso_inv, + Presheaf.tautologicalCocone_ι_app, HasColimit.isoOfNatIso_ι_hom_assoc, + HasLimit.isoOfNatIso_hom_π_assoc, HasColimit.isoOfNatIso_ι_hom_assoc, + HasColimit.isoOfNatIso_ι_hom_assoc, HasColimit.isoOfNatIso_ι_hom_assoc, + ι_colimitLimitIso_limit_π_assoc, isoAux_hom_app, ← NatTrans.comp_app_assoc, + ← NatTrans.comp_app_assoc, Category.assoc, HasLimit.isoOfNatIso_hom_π, + preservesLimitIso_hom_π_assoc, Iso.symm_hom, + ← NatTrans.comp_app_assoc, HasColimit.isoOfNatIso_ι_hom, + ← NatTrans.comp_app_assoc, Category.assoc, + ι_colimitCompWhiskeringLeftIsoCompColimit_hom, + NatTrans.comp_app, Category.assoc, isoWhiskerLeft_hom, NatTrans.comp_app, Category.assoc, + ← NatTrans.comp_app, ← whiskerLeft_comp, colimit.comp_coconePointUniqueUpToIso_hom] + + have := i.hom.naturality (limit.π K j) + dsimp only [yoneda_obj_obj, Functor.const_obj_obj] at this + rw [← this] + ext + simp + +theorem isIso_post [IsFiltered (CostructuredArrow yoneda A)] : IsIso (limit.post K A) := + iso_hom A K ▸ inferInstance + +end PreservesFiniteLimitsOfIsFilteredCostructuredArrowYonedaAux + +attribute [local instance] PreservesFiniteLimitsOfIsFilteredCostructuredArrowYonedaAux.isIso_post + +/-- If `C` is a small finitely cocomplete category and `A : Cᵒᵖ ⥤ Type u` is a presheaf such that +`CostructuredArrow yoneda A` is filtered, then `A` preserves finite limits. + +One direction of Proposition 3.3.13 of [Kashiwara2006]. +-/ +noncomputable def preservesFiniteLimitsOfIsFilteredCostructuredArrowYoneda + [IsFiltered (CostructuredArrow yoneda A)] : PreservesFiniteLimits A where + preservesFiniteLimits _ _ _ := ⟨fun {_} => preservesLimitOfIsIsoPost _ _⟩ + +/-- If `C` is a small finitely cocomplete category and `A : Cᵒᵖ ⥤ Type u` is a presheaf, then +`CostructuredArrow yoneda A` is filtered if and only if `A` preserves finite limits. + +Proposition 3.3.13 of [Kashiwara2006]. +-/ +theorem isFiltered_costructuredArrow_yoneda_iff_nonempty_preservesFiniteLimits : + IsFiltered (CostructuredArrow yoneda A) ↔ Nonempty (PreservesFiniteLimits A) := + ⟨fun _ => ⟨preservesFiniteLimitsOfIsFilteredCostructuredArrowYoneda A⟩, + fun ⟨_⟩ => isFiltered_costructuredArrow_yoneda_of_preservesFiniteLimits A⟩ + +end CategoryTheory.Limits From 8bf29506f05e62812a25153bcf760afdac48ce8d Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 18 Nov 2024 18:54:16 +0000 Subject: [PATCH 084/829] ci: ensure error codes from `git rev-parse` are propagated (#19215) Previously the exit status of the git command would not be propagated --- .github/workflows/update_dependencies.yml | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/.github/workflows/update_dependencies.yml b/.github/workflows/update_dependencies.yml index 27646aaec561b..48111dc1fac1d 100644 --- a/.github/workflows/update_dependencies.yml +++ b/.github/workflows/update_dependencies.yml @@ -25,7 +25,8 @@ jobs: - name: Get sha of branch id: sha run: | - echo "sha=$(git rev-parse --verify origin/update-dependencies-bot-use-only)" >> "$GITHUB_OUTPUT" + SHA="$(git rev-parse --verify origin/update-dependencies-bot-use-only)" + echo "sha=$SHA" >> "$GITHUB_OUTPUT" - name: Get PR and labels if: ${{ steps.sha.outputs.sha }} From c34ded5a563a8eea9cf7be877642b46dac377a85 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <425260+eric-wieser@users.noreply.github.com> Date: Mon, 18 Nov 2024 18:54:17 +0000 Subject: [PATCH 085/829] chore: update Mathlib dependencies 2024-11-18 (#19219) 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 ae0dc1f2909f6..2464526042495 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "b100ff2565805e9f30a482788b3fc66937a7f38a", + "rev": "01f4969b6e861db6a99261ea5eadd5a9bb63011b", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", @@ -55,7 +55,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "b0b73e5bc33f1bc4d3c0f254630dd0e262cecc08", + "rev": "119b022b3ea88ec810a677888528e50f8144a26e", "name": "importGraph", "manifestFile": "lake-manifest.json", "inputRev": "main", From efe2a970a4e4f87559b5e56166f7b5324973b661 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 18 Nov 2024 20:53:41 +0000 Subject: [PATCH 086/829] =?UTF-8?q?feat(MvPolynomial):=20`C=20a=20=3D=200?= =?UTF-8?q?=20=E2=86=94=20a=20=3D=200`=20(#19202)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/Algebra/MonoidAlgebra/Degree.lean | 3 +++ Mathlib/Algebra/MvPolynomial/Basic.lean | 5 +++++ 2 files changed, 8 insertions(+) diff --git a/Mathlib/Algebra/MonoidAlgebra/Degree.lean b/Mathlib/Algebra/MonoidAlgebra/Degree.lean index 16c11daef0708..55345300a014e 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Degree.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Degree.lean @@ -411,6 +411,9 @@ lemma leadingCoeff_eq_zero (hD : D.Injective) : p.leadingCoeff D = 0 ↔ p = 0 : rw [leadingCoeff, ← Ne, ← Finsupp.mem_support_iff] exact supDegree_mem_support hD h +lemma leadingCoeff_ne_zero (hD : D.Injective) : p.leadingCoeff D ≠ 0 ↔ p ≠ 0 := + (leadingCoeff_eq_zero hD).ne + lemma supDegree_sub_lt_of_leadingCoeff_eq (hD : D.Injective) {R} [CommRing R] {p q : R[A]} (hd : p.supDegree D = q.supDegree D) (hc : p.leadingCoeff D = q.leadingCoeff D) : (p - q).supDegree D < p.supDegree D ∨ p = q := by diff --git a/Mathlib/Algebra/MvPolynomial/Basic.lean b/Mathlib/Algebra/MvPolynomial/Basic.lean index f1c838c583594..dac73ae71fe7a 100644 --- a/Mathlib/Algebra/MvPolynomial/Basic.lean +++ b/Mathlib/Algebra/MvPolynomial/Basic.lean @@ -231,6 +231,11 @@ theorem C_inj {σ : Type*} (R : Type*) [CommSemiring R] (r s : R) : (C r : MvPolynomial σ R) = C s ↔ r = s := (C_injective σ R).eq_iff +@[simp] lemma C_eq_zero : (C a : MvPolynomial σ R) = 0 ↔ a = 0 := by rw [← map_zero C, C_inj] + +lemma C_ne_zero : (C a : MvPolynomial σ R) ≠ 0 ↔ a ≠ 0 := + C_eq_zero.ne + instance nontrivial_of_nontrivial (σ : Type*) (R : Type*) [CommSemiring R] [Nontrivial R] : Nontrivial (MvPolynomial σ R) := inferInstanceAs (Nontrivial <| AddMonoidAlgebra R (σ →₀ ℕ)) From 654802e5ec7c62e0cf1f45dfbdaa746af1b81967 Mon Sep 17 00:00:00 2001 From: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com> Date: Mon, 18 Nov 2024 21:18:06 +0000 Subject: [PATCH 087/829] feat(Polynomial/Splits): dvd_iff_roots_le_roots (#18800) Co-authored-by: leanprover-community-mathlib4-bot Co-authored-by: ChrisHughes24 --- Mathlib/Algebra/Polynomial/Splits.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/Mathlib/Algebra/Polynomial/Splits.lean b/Mathlib/Algebra/Polynomial/Splits.lean index db9df3da20b50..98d5a93538fdf 100644 --- a/Mathlib/Algebra/Polynomial/Splits.lean +++ b/Mathlib/Algebra/Polynomial/Splits.lean @@ -322,6 +322,18 @@ theorem eq_prod_roots_of_splits_id {p : K[X]} (hsplit : Splits (RingHom.id K) p) p = C p.leadingCoeff * (p.roots.map fun a => X - C a).prod := by simpa using eq_prod_roots_of_splits hsplit +theorem Splits.dvd_of_roots_le_roots {p q : K[X]} (hp : p.Splits (RingHom.id _)) (hp0 : p ≠ 0) + (hq : p.roots ≤ q.roots) : p ∣ q := by + rw [eq_prod_roots_of_splits_id hp, C_mul_dvd (leadingCoeff_ne_zero.2 hp0)] + exact dvd_trans + (Multiset.prod_dvd_prod_of_le (Multiset.map_le_map hq)) + (prod_multiset_X_sub_C_dvd _) + +theorem Splits.dvd_iff_roots_le_roots {p q : K[X]} + (hp : p.Splits (RingHom.id _)) (hp0 : p ≠ 0) (hq0 : q ≠ 0) : + p ∣ q ↔ p.roots ≤ q.roots := + ⟨Polynomial.roots.le_of_dvd hq0, hp.dvd_of_roots_le_roots hp0⟩ + theorem aeval_eq_prod_aroots_sub_of_splits [Algebra K L] {p : K[X]} (hsplit : Splits (algebraMap K L) p) (v : L) : aeval v p = algebraMap K L p.leadingCoeff * ((p.aroots L).map fun a ↦ v - a).prod := by From dc324b56f7527bc59d5fcf33bf1c7f73a8c91d9d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 18 Nov 2024 21:32:21 +0000 Subject: [PATCH 088/829] feat: right-division as an `OrderIso` (#19191) Co-authored-by: Andrew Yang --- .../Algebra/GroupWithZero/Units/Equiv.lean | 21 +++++++++++++------ .../Order/GroupWithZero/Unbundled/Lemmas.lean | 10 +++++++++ 2 files changed, 25 insertions(+), 6 deletions(-) diff --git a/Mathlib/Algebra/GroupWithZero/Units/Equiv.lean b/Mathlib/Algebra/GroupWithZero/Units/Equiv.lean index bc6167e7b9547..9dda8f77c9de9 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Equiv.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Equiv.lean @@ -12,32 +12,41 @@ import Mathlib.Algebra.GroupWithZero.Units.Basic assert_not_exists DenselyOrdered -variable {G : Type*} +variable {G₀ : Type*} namespace Equiv section GroupWithZero -variable [GroupWithZero G] +variable [GroupWithZero G₀] /-- Left multiplication by a nonzero element in a `GroupWithZero` is a permutation of the underlying type. -/ @[simps! (config := .asFn)] -protected def mulLeft₀ (a : G) (ha : a ≠ 0) : Perm G := +protected def mulLeft₀ (a : G₀) (ha : a ≠ 0) : Perm G₀ := (Units.mk0 a ha).mulLeft -theorem _root_.mulLeft_bijective₀ (a : G) (ha : a ≠ 0) : Function.Bijective (a * · : G → G) := +theorem _root_.mulLeft_bijective₀ (a : G₀) (ha : a ≠ 0) : Function.Bijective (a * · : G₀ → G₀) := (Equiv.mulLeft₀ a ha).bijective /-- Right multiplication by a nonzero element in a `GroupWithZero` is a permutation of the underlying type. -/ @[simps! (config := .asFn)] -protected def mulRight₀ (a : G) (ha : a ≠ 0) : Perm G := +protected def mulRight₀ (a : G₀) (ha : a ≠ 0) : Perm G₀ := (Units.mk0 a ha).mulRight -theorem _root_.mulRight_bijective₀ (a : G) (ha : a ≠ 0) : Function.Bijective ((· * a) : G → G) := +theorem _root_.mulRight_bijective₀ (a : G₀) (ha : a ≠ 0) : Function.Bijective ((· * a) : G₀ → G₀) := (Equiv.mulRight₀ a ha).bijective +/-- Right division by a nonzero element in a `GroupWithZero` is a permutation of the +underlying type. -/ +@[simps! (config := { simpRhs := true })] +def divRight₀ (a : G₀) (ha : a ≠ 0) : G₀ ≃ G₀ where + toFun := (· / a) + invFun := (· * a) + left_inv _ := by simp [ha] + right_inv _ := by simp [ha] + end GroupWithZero end Equiv diff --git a/Mathlib/Algebra/Order/GroupWithZero/Unbundled/Lemmas.lean b/Mathlib/Algebra/Order/GroupWithZero/Unbundled/Lemmas.lean index 53d77b1ec9d19..15331ca9c673a 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Unbundled/Lemmas.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Unbundled/Lemmas.lean @@ -25,3 +25,13 @@ def OrderIso.mulLeft₀ [PosMulMono G₀] [PosMulReflectLE G₀] (a : G₀) (ha 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 + +/-- `Equiv.divRight₀` as an order isomorphism. -/ +@[simps! (config := { simpRhs := true })] +def OrderIso.divRight₀ {G₀} [GroupWithZero G₀] [PartialOrder G₀] [ZeroLEOneClass G₀] + [MulPosStrictMono G₀] [MulPosReflectLE G₀] [PosMulReflectLT G₀] + (a : G₀) (ha : 0 < a) : G₀ ≃o G₀ where + toEquiv := .divRight₀ a ha.ne' + map_rel_iff' {b c} := by + simp only [Equiv.divRight₀_apply, div_eq_mul_inv] + exact mul_le_mul_right (a := a⁻¹) (inv_pos.mpr ha) From 436211744e3cfc6ceefe33edf3f22adf4fce6aff Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Mon, 18 Nov 2024 21:32:24 +0000 Subject: [PATCH 089/829] feat(CategoryTheory): weaken assumptions and dualize file `Limits/Constructions/Filtered.lean` (#19196) --- .../Limits/Constructions/Filtered.lean | 78 +++++++++++++++++-- 1 file changed, 73 insertions(+), 5 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/Constructions/Filtered.lean b/Mathlib/CategoryTheory/Limits/Constructions/Filtered.lean index e7e1a5ed68f8b..5712931af86a2 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/Filtered.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/Filtered.lean @@ -46,8 +46,8 @@ def liftToFinsetObj (F : Discrete α ⥤ C) : Finset (Discrete α) ⥤ C where taking the colimit of the diagram formed by the coproducts of finite sets over the indexing type. -/ @[simps!] -def liftToFinsetColimitCocone [HasFilteredColimitsOfSize.{w, w} C] (F : Discrete α ⥤ C) : - ColimitCocone F where +def liftToFinsetColimitCocone [HasColimitsOfShape (Finset (Discrete α)) C] + (F : Discrete α ⥤ C) : ColimitCocone F where cocone := { pt := colimit (liftToFinsetObj F) ι := @@ -106,9 +106,8 @@ namespace CoproductsFromFiniteFiltered section -variable [HasFiniteCoproducts C] [HasFilteredColimitsOfSize.{w, w} C] - -attribute [local instance] hasCoproducts_of_finite_and_filtered +variable [HasFiniteCoproducts C] [HasColimitsOfShape (Finset (Discrete α)) C] + [HasColimitsOfShape (Discrete α) C] /-- Helper construction for `liftToFinsetColimIso`. -/ @[reassoc] @@ -144,4 +143,73 @@ def liftToFinsetEvaluationIso [HasFiniteCoproducts C] (I : Finset (Discrete α)) end CoproductsFromFiniteFiltered +namespace ProductsFromFiniteCofiltered + +variable [HasFiniteProducts C] + +/-- If `C` has finite coproducts, a functor `Discrete α ⥤ C` lifts to a functor + `Finset (Discrete α) ⥤ C` by taking coproducts. -/ +@[simps!] +def liftToFinsetObj (F : Discrete α ⥤ C) : (Finset (Discrete α))ᵒᵖ ⥤ C where + obj s := ∏ᶜ (fun x : s.unop => F.obj x) + map {Y _} h := Pi.lift fun y => + Pi.π (fun (x : { x // x ∈ Y.unop }) => F.obj x) ⟨y, h.unop.down.down y.2⟩ + + +/-- If `C` has finite coproducts and filtered colimits, we can construct arbitrary coproducts by + taking the colimit of the diagram formed by the coproducts of finite sets over the indexing + type. -/ +@[simps!] +def liftToFinsetLimitCone [HasLimitsOfShape (Finset (Discrete α))ᵒᵖ C] + (F : Discrete α ⥤ C) : LimitCone F where + cone := + { pt := limit (liftToFinsetObj F) + π := Discrete.natTrans fun j => + limit.π (liftToFinsetObj F) ⟨{j}⟩ ≫ Pi.π _ (⟨j, by simp⟩ : ({j} : Finset (Discrete α))) } + isLimit := + { lift := fun s => + limit.lift (liftToFinsetObj F) + { pt := s.pt + π := { app := fun _ => Pi.lift fun x => s.π.app x } } + uniq := fun s m h => by + apply limit.hom_ext + rintro t + dsimp [liftToFinsetObj] + apply limit.hom_ext + rintro ⟨⟨j, hj⟩⟩ + convert h j using 1 + · simp [← limit.w (liftToFinsetObj F) ⟨⟨⟨Finset.singleton_subset_iff.2 hj⟩⟩⟩] + rfl + · aesop_cat } + +variable (C) (α) + +/-- The functor taking a functor `Discrete α ⥤ C` to a functor `Finset (Discrete α) ⥤ C` by taking +coproducts. -/ +@[simps!] +def liftToFinset : (Discrete α ⥤ C) ⥤ ((Finset (Discrete α))ᵒᵖ ⥤ C) where + obj := liftToFinsetObj + map := fun β => { app := fun _ => Pi.map (fun x => β.app x.val) } + +/-- The `liftToFinset` functor, precomposed with forming a colimit, is a coproduct on the original +functor. -/ +def liftToFinsetLimIso [HasLimitsOfShape (Finset (Discrete α))ᵒᵖ C] + [HasLimitsOfShape (Discrete α) C] : liftToFinset C α ⋙ lim ≅ lim := + NatIso.ofComponents + (fun F => Iso.symm <| limit.isoLimitCone (liftToFinsetLimitCone F)) + (fun β => by + simp only [Functor.comp_obj, lim_obj, Functor.comp_map, lim_map, Iso.symm_hom] + ext J + simp [liftToFinset]) + +/-- `liftToFinset`, when composed with the evaluation functor, results in the whiskering composed +with `colim`. -/ +def liftToFinsetEvaluationIso (I : Finset (Discrete α)) : + liftToFinset C α ⋙ (evaluation _ _).obj ⟨I⟩ ≅ + (whiskeringLeft _ _ _).obj (Discrete.functor (·.val)) ⋙ lim (J := Discrete I) := + NatIso.ofComponents (fun _ => HasLimit.isoOfNatIso (Discrete.natIso fun _ => Iso.refl _)) + fun _ => by dsimp; ext; simp + +end ProductsFromFiniteCofiltered + end CategoryTheory.Limits From db980d623801cca71c769d5eee52c9cc8e3cfa63 Mon Sep 17 00:00:00 2001 From: FR Date: Mon, 18 Nov 2024 22:39:35 +0000 Subject: [PATCH 090/829] =?UTF-8?q?feat:=20define=20`PGame.identical`=20`P?= =?UTF-8?q?Game.mem=E2=82=97`=20`PGame.mem=E1=B5=A3`=20(#17122)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/SetTheory/Game/PGame.lean | 146 +++++++++++++++++++++++++++++- 1 file changed, 144 insertions(+), 2 deletions(-) diff --git a/Mathlib/SetTheory/Game/PGame.lean b/Mathlib/SetTheory/Game/PGame.lean index 50576e07cbe0b..63640f6a65f95 100644 --- a/Mathlib/SetTheory/Game/PGame.lean +++ b/Mathlib/SetTheory/Game/PGame.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2019 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Reid Barton, Mario Carneiro, Isabel Longbottom, Kim Morrison +Authors: Reid Barton, Mario Carneiro, Isabel Longbottom, Kim Morrison, Yuyang Zhao -/ import Mathlib.Algebra.Order.ZeroLEOne import Mathlib.Data.List.InsertIdx @@ -343,6 +343,139 @@ instance uniqueOneLeftMoves : Unique (LeftMoves 1) := instance isEmpty_one_rightMoves : IsEmpty (RightMoves 1) := instIsEmptyPEmpty +/-! ### Identity -/ + +/-- Two pre-games are identical if their left and right sets are identical. +That is, `Identical x y` if every left move of `x` is identical to some left move of `y`, +every right move of `x` is identical to some right move of `y`, and vice versa. -/ +def Identical : PGame.{u} → PGame.{u} → Prop + | mk _ _ xL xR, mk _ _ yL yR => + Relator.BiTotal (fun i j ↦ Identical (xL i) (yL j)) ∧ + Relator.BiTotal (fun i j ↦ Identical (xR i) (yR j)) + +@[inherit_doc] scoped infix:50 " ≡ " => PGame.Identical + +theorem identical_iff : ∀ {x y : PGame}, x ≡ y ↔ + Relator.BiTotal (x.moveLeft · ≡ y.moveLeft ·) ∧ Relator.BiTotal (x.moveRight · ≡ y.moveRight ·) + | mk _ _ _ _, mk _ _ _ _ => Iff.rfl + +@[refl, simp] protected theorem Identical.refl (x) : x ≡ x := + PGame.recOn x fun _ _ _ _ IHL IHR ↦ ⟨Relator.BiTotal.refl IHL, Relator.BiTotal.refl IHR⟩ + +protected theorem Identical.rfl {x} : x ≡ x := Identical.refl x + +@[symm] protected theorem Identical.symm : ∀ {x y}, x ≡ y → y ≡ x + | mk _ _ _ _, mk _ _ _ _, ⟨hL, hR⟩ => ⟨hL.symm fun _ _ h ↦ h.symm, hR.symm fun _ _ h ↦ h.symm⟩ + +theorem identical_comm {x y} : x ≡ y ↔ y ≡ x := + ⟨.symm, .symm⟩ + +@[trans] protected theorem Identical.trans : ∀ {x y z}, x ≡ y → y ≡ z → x ≡ z + | mk _ _ _ _, mk _ _ _ _, mk _ _ _ _, ⟨hL₁, hR₁⟩, ⟨hL₂, hR₂⟩ => + ⟨hL₁.trans (fun _ _ _ h₁ h₂ ↦ h₁.trans h₂) hL₂, hR₁.trans (fun _ _ _ h₁ h₂ ↦ h₁.trans h₂) hR₂⟩ + +/-- `x ∈ₗ y` if `x` is identical to some left move of `y`. -/ +def memₗ (x y : PGame.{u}) : Prop := ∃ b, x ≡ y.moveLeft b + +/-- `x ∈ᵣ y` if `x` is identical to some right move of `y`. -/ +def memᵣ (x y : PGame.{u}) : Prop := ∃ b, x ≡ y.moveRight b + +@[inherit_doc] scoped infix:50 " ∈ₗ " => PGame.memₗ +@[inherit_doc] scoped infix:50 " ∈ᵣ " => PGame.memᵣ +@[inherit_doc PGame.memₗ] binder_predicate x " ∈ₗ " y:term => `($x ∈ₗ $y) +@[inherit_doc PGame.memᵣ] binder_predicate x " ∈ᵣ " y:term => `($x ∈ᵣ $y) + +theorem memₗ_def {x y : PGame} : x ∈ₗ y ↔ ∃ b, x ≡ y.moveLeft b := .rfl +theorem memᵣ_def {x y : PGame} : x ∈ᵣ y ↔ ∃ b, x ≡ y.moveRight b := .rfl +theorem moveLeft_memₗ (x : PGame) (b) : x.moveLeft b ∈ₗ x := ⟨_, .rfl⟩ +theorem moveRight_memᵣ (x : PGame) (b) : x.moveRight b ∈ᵣ x := ⟨_, .rfl⟩ + +theorem identical_of_isEmpty (x y : PGame) + [IsEmpty x.LeftMoves] [IsEmpty x.RightMoves] + [IsEmpty y.LeftMoves] [IsEmpty y.RightMoves] : x ≡ y := + identical_iff.2 <| by simp [Relator.BiTotal, Relator.LeftTotal, Relator.RightTotal] + +/-- `Identical` as a `Setoid`. -/ +def identicalSetoid : Setoid PGame := + ⟨Identical, Identical.refl, Identical.symm, Identical.trans⟩ + +instance : IsRefl PGame (· ≡ ·) := ⟨Identical.refl⟩ +instance : IsSymm PGame (· ≡ ·) := ⟨fun _ _ ↦ Identical.symm⟩ +instance : IsTrans PGame (· ≡ ·) := ⟨fun _ _ _ ↦ Identical.trans⟩ +instance : IsEquiv PGame (· ≡ ·) := { } + +/-- If `x` and `y` are identical, then a left move of `x` is identical to some left move of `y`. -/ +lemma Identical.moveLeft : ∀ {x y}, x ≡ y → + ∀ i, ∃ j, x.moveLeft i ≡ y.moveLeft j + | mk _ _ _ _, mk _ _ _ _, ⟨hl, _⟩, i => hl.1 i + +/-- If `x` and `y` are identical, then a right move of `x` is identical to some right move of `y`. +-/ +lemma Identical.moveRight : ∀ {x y}, x ≡ y → + ∀ i, ∃ j, x.moveRight i ≡ y.moveRight j + | mk _ _ _ _, mk _ _ _ _, ⟨_, hr⟩, i => hr.1 i + +/-- Uses `∈ₗ` and `∈ᵣ` instead of `≡`. -/ +theorem identical_iff' : ∀ {x y : PGame}, x ≡ y ↔ + ((∀ i, x.moveLeft i ∈ₗ y) ∧ (∀ j, y.moveLeft j ∈ₗ x)) ∧ + ((∀ i, x.moveRight i ∈ᵣ y) ∧ (∀ j, y.moveRight j ∈ᵣ x)) + | mk xl xr xL xR, mk yl yr yL yR => by + convert identical_iff <;> + dsimp [Relator.BiTotal, Relator.LeftTotal, Relator.RightTotal] <;> + congr! <;> + exact exists_congr <| fun _ ↦ identical_comm + +theorem memₗ.congr_right : ∀ {x y : PGame}, + x ≡ y → (∀ {w : PGame}, w ∈ₗ x ↔ w ∈ₗ y) + | mk _ _ _ _, mk _ _ _ _, ⟨⟨h₁, h₂⟩, _⟩, _w => + ⟨fun ⟨i, hi⟩ ↦ (h₁ i).imp (fun _ ↦ hi.trans), + fun ⟨j, hj⟩ ↦ (h₂ j).imp (fun _ hi ↦ hj.trans hi.symm)⟩ + +theorem memᵣ.congr_right : ∀ {x y : PGame}, + x ≡ y → (∀ {w : PGame}, w ∈ᵣ x ↔ w ∈ᵣ y) + | mk _ _ _ _, mk _ _ _ _, ⟨_, ⟨h₁, h₂⟩⟩, _w => + ⟨fun ⟨i, hi⟩ ↦ (h₁ i).imp (fun _ ↦ hi.trans), + fun ⟨j, hj⟩ ↦ (h₂ j).imp (fun _ hi ↦ hj.trans hi.symm)⟩ + +theorem memₗ.congr_left : ∀ {x y : PGame}, + x ≡ y → (∀ {w : PGame}, x ∈ₗ w ↔ y ∈ₗ w) + | _, _, h, mk _ _ _ _ => ⟨fun ⟨i, hi⟩ ↦ ⟨i, h.symm.trans hi⟩, fun ⟨i, hi⟩ ↦ ⟨i, h.trans hi⟩⟩ + +theorem memᵣ.congr_left : ∀ {x y : PGame}, + x ≡ y → (∀ {w : PGame}, x ∈ᵣ w ↔ y ∈ᵣ w) + | _, _, h, mk _ _ _ _ => ⟨fun ⟨i, hi⟩ ↦ ⟨i, h.symm.trans hi⟩, fun ⟨i, hi⟩ ↦ ⟨i, h.trans hi⟩⟩ + +lemma Identical.ext : ∀ {x y}, (∀ z, z ∈ₗ x ↔ z ∈ₗ y) → (∀ z, z ∈ᵣ x ↔ z ∈ᵣ y) → x ≡ y + | mk _ _ _ _, mk _ _ _ _, hl, hr => identical_iff'.mpr + ⟨⟨fun i ↦ (hl _).mp ⟨i, refl _⟩, fun j ↦ (hl _).mpr ⟨j, refl _⟩⟩, + ⟨fun i ↦ (hr _).mp ⟨i, refl _⟩, fun j ↦ (hr _).mpr ⟨j, refl _⟩⟩⟩ + +lemma Identical.ext_iff {x y} : x ≡ y ↔ (∀ z, z ∈ₗ x ↔ z ∈ₗ y) ∧ (∀ z, z ∈ᵣ x ↔ z ∈ᵣ y) := + ⟨fun h ↦ ⟨@memₗ.congr_right _ _ h, @memᵣ.congr_right _ _ h⟩, fun h ↦ h.elim Identical.ext⟩ + +lemma Identical.congr_right {x y z} (h : x ≡ y) : z ≡ x ↔ z ≡ y := + ⟨fun hz ↦ hz.trans h, fun hz ↦ hz.trans h.symm⟩ + +lemma Identical.congr_left {x y z} (h : x ≡ y) : x ≡ z ↔ y ≡ z := + ⟨fun hz ↦ h.symm.trans hz, fun hz ↦ h.trans hz⟩ + +/-- Show `x ≡ y` by giving an explicit correspondence between the moves of `x` and `y`. -/ +lemma Identical.of_fn {x y : PGame} + (l : x.LeftMoves → y.LeftMoves) (il : y.LeftMoves → x.LeftMoves) + (r : x.RightMoves → y.RightMoves) (ir : y.RightMoves → x.RightMoves) + (hl : ∀ i, x.moveLeft i ≡ y.moveLeft (l i)) + (hil : ∀ i, x.moveLeft (il i) ≡ y.moveLeft i) + (hr : ∀ i, x.moveRight i ≡ y.moveRight (r i)) + (hir : ∀ i, x.moveRight (ir i) ≡ y.moveRight i) : x ≡ y := + identical_iff.mpr + ⟨⟨fun i ↦ ⟨l i, hl i⟩, fun i ↦ ⟨il i, hil i⟩⟩, ⟨fun i ↦ ⟨r i, hr i⟩, fun i ↦ ⟨ir i, hir i⟩⟩⟩ + +lemma Identical.of_equiv {x y : PGame} + (l : x.LeftMoves ≃ y.LeftMoves) (r : x.RightMoves ≃ y.RightMoves) + (hl : ∀ i, x.moveLeft i ≡ y.moveLeft (l i)) (hr : ∀ i, x.moveRight i ≡ y.moveRight (r i)) : + x ≡ y := + .of_fn l l.symm r r.symm hl (by simpa using hl <| l.symm ·) hr (by simpa using hr <| r.symm ·) + /-! ### Pre-game order relations -/ @@ -477,6 +610,13 @@ instance : Preorder PGame := le_trans_aux (fun {i} => (IHzl i).1) fun {j} => (IHyr j).2.1⟩ lt := fun x y => x ≤ y ∧ x ⧏ y } +lemma Identical.le : ∀ {x y}, x ≡ y → x ≤ y + | mk _ _ _ _, mk _ _ _ _, ⟨hL, hR⟩ => le_of_forall_lf + (fun i ↦ let ⟨_, hj⟩ := hL.1 i; lf_of_le_moveLeft hj.le) + (fun i ↦ let ⟨_, hj⟩ := hR.2 i; lf_of_moveRight_le hj.le) + +lemma Identical.ge {x y} (h : x ≡ y) : y ≤ x := h.symm.le + theorem lt_iff_le_and_lf {x y : PGame} : x < y ↔ x ≤ y ∧ x ⧏ y := Iff.rfl @@ -716,6 +856,8 @@ protected theorem equiv_comm {x y : PGame} : (x ≈ y) ↔ (y ≈ x) := theorem equiv_of_eq {x y : PGame} (h : x = y) : x ≈ y := by subst h; rfl +lemma Identical.equiv {x y} (h : x ≡ y) : x ≈ y := ⟨h.le, h.ge⟩ + @[trans] theorem le_of_le_of_equiv {x y z : PGame} (h₁ : x ≤ y) (h₂ : y ≈ z) : x ≤ z := h₁.trans h₂.1 @@ -1795,4 +1937,4 @@ end PGame end SetTheory -set_option linter.style.longFile 1900 +set_option linter.style.longFile 2100 From 7c210ffe4181188eb7be12bcb4297e0b5659808d Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Mon, 18 Nov 2024 23:38:29 +0000 Subject: [PATCH 091/829] chore: move NumberTheory.Liouville to NumberTheory.Transcendental.Liouville (#19225) In preparation for #6718. --- Mathlib.lean | 10 +++++----- .../{ => Transcendental}/Liouville/Basic.lean | 0 .../Liouville/LiouvilleNumber.lean | 2 +- .../{ => Transcendental}/Liouville/LiouvilleWith.lean | 2 +- .../{ => Transcendental}/Liouville/Measure.lean | 4 ++-- .../{ => Transcendental}/Liouville/Residual.lean | 2 +- 6 files changed, 10 insertions(+), 10 deletions(-) rename Mathlib/NumberTheory/{ => Transcendental}/Liouville/Basic.lean (100%) rename Mathlib/NumberTheory/{ => Transcendental}/Liouville/LiouvilleNumber.lean (99%) rename Mathlib/NumberTheory/{ => Transcendental}/Liouville/LiouvilleWith.lean (99%) rename Mathlib/NumberTheory/{ => Transcendental}/Liouville/Measure.lean (97%) rename Mathlib/NumberTheory/{ => Transcendental}/Liouville/Residual.lean (98%) diff --git a/Mathlib.lean b/Mathlib.lean index 2314aad484f69..e9bea276ba9f6 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3766,11 +3766,6 @@ import Mathlib.NumberTheory.LegendreSymbol.QuadraticChar.Basic import Mathlib.NumberTheory.LegendreSymbol.QuadraticChar.GaussSum import Mathlib.NumberTheory.LegendreSymbol.QuadraticReciprocity import Mathlib.NumberTheory.LegendreSymbol.ZModChar -import Mathlib.NumberTheory.Liouville.Basic -import Mathlib.NumberTheory.Liouville.LiouvilleNumber -import Mathlib.NumberTheory.Liouville.LiouvilleWith -import Mathlib.NumberTheory.Liouville.Measure -import Mathlib.NumberTheory.Liouville.Residual import Mathlib.NumberTheory.LucasLehmer import Mathlib.NumberTheory.LucasPrimality import Mathlib.NumberTheory.MaricaSchoenheim @@ -3834,6 +3829,11 @@ import Mathlib.NumberTheory.SmoothNumbers import Mathlib.NumberTheory.SumFourSquares import Mathlib.NumberTheory.SumPrimeReciprocals import Mathlib.NumberTheory.SumTwoSquares +import Mathlib.NumberTheory.Transcendental.Liouville.Basic +import Mathlib.NumberTheory.Transcendental.Liouville.LiouvilleNumber +import Mathlib.NumberTheory.Transcendental.Liouville.LiouvilleWith +import Mathlib.NumberTheory.Transcendental.Liouville.Measure +import Mathlib.NumberTheory.Transcendental.Liouville.Residual import Mathlib.NumberTheory.VonMangoldt import Mathlib.NumberTheory.WellApproximable import Mathlib.NumberTheory.Wilson diff --git a/Mathlib/NumberTheory/Liouville/Basic.lean b/Mathlib/NumberTheory/Transcendental/Liouville/Basic.lean similarity index 100% rename from Mathlib/NumberTheory/Liouville/Basic.lean rename to Mathlib/NumberTheory/Transcendental/Liouville/Basic.lean diff --git a/Mathlib/NumberTheory/Liouville/LiouvilleNumber.lean b/Mathlib/NumberTheory/Transcendental/Liouville/LiouvilleNumber.lean similarity index 99% rename from Mathlib/NumberTheory/Liouville/LiouvilleNumber.lean rename to Mathlib/NumberTheory/Transcendental/Liouville/LiouvilleNumber.lean index 093de3c611294..72b1c28fd82b7 100644 --- a/Mathlib/NumberTheory/Liouville/LiouvilleNumber.lean +++ b/Mathlib/NumberTheory/Transcendental/Liouville/LiouvilleNumber.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Jujian Zhang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa, Jujian Zhang -/ -import Mathlib.NumberTheory.Liouville.Basic +import Mathlib.NumberTheory.Transcendental.Liouville.Basic /-! diff --git a/Mathlib/NumberTheory/Liouville/LiouvilleWith.lean b/Mathlib/NumberTheory/Transcendental/Liouville/LiouvilleWith.lean similarity index 99% rename from Mathlib/NumberTheory/Liouville/LiouvilleWith.lean rename to Mathlib/NumberTheory/Transcendental/Liouville/LiouvilleWith.lean index 8406047fad797..6e1ac10afa4e9 100644 --- a/Mathlib/NumberTheory/Liouville/LiouvilleWith.lean +++ b/Mathlib/NumberTheory/Transcendental/Liouville/LiouvilleWith.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Analysis.SpecialFunctions.Pow.Asymptotics -import Mathlib.NumberTheory.Liouville.Basic +import Mathlib.NumberTheory.Transcendental.Liouville.Basic import Mathlib.Topology.Instances.Irrational /-! diff --git a/Mathlib/NumberTheory/Liouville/Measure.lean b/Mathlib/NumberTheory/Transcendental/Liouville/Measure.lean similarity index 97% rename from Mathlib/NumberTheory/Liouville/Measure.lean rename to Mathlib/NumberTheory/Transcendental/Liouville/Measure.lean index 55c85ab7d9a3d..ca97dfa42b9e9 100644 --- a/Mathlib/NumberTheory/Liouville/Measure.lean +++ b/Mathlib/NumberTheory/Transcendental/Liouville/Measure.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.MeasureTheory.Measure.Lebesgue.Basic -import Mathlib.NumberTheory.Liouville.Residual -import Mathlib.NumberTheory.Liouville.LiouvilleWith +import Mathlib.NumberTheory.Transcendental.Liouville.Residual +import Mathlib.NumberTheory.Transcendental.Liouville.LiouvilleWith import Mathlib.Analysis.PSeries /-! diff --git a/Mathlib/NumberTheory/Liouville/Residual.lean b/Mathlib/NumberTheory/Transcendental/Liouville/Residual.lean similarity index 98% rename from Mathlib/NumberTheory/Liouville/Residual.lean rename to Mathlib/NumberTheory/Transcendental/Liouville/Residual.lean index 151dd1af3c8f8..48892d46d0103 100644 --- a/Mathlib/NumberTheory/Liouville/Residual.lean +++ b/Mathlib/NumberTheory/Transcendental/Liouville/Residual.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.NumberTheory.Liouville.Basic +import Mathlib.NumberTheory.Transcendental.Liouville.Basic import Mathlib.Topology.Baire.Lemmas import Mathlib.Topology.Baire.LocallyCompactRegular import Mathlib.Topology.Instances.Irrational From 42f07b979f97da335989d3a48ecaf4f0d46ba688 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 19 Nov 2024 00:05:26 +0000 Subject: [PATCH 092/829] refactor(MeasureTheory/MeasurableSpace/Card): avoid `Ordinal.toType` (#18199) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We redefine `generateMeasurableRec` so that it's indexed by an ordinal, rather than being indexed by `ω₁.toType`. We also employ `ω₁` and its new API throughout the file. --- .../MeasureTheory/MeasurableSpace/Card.lean | 223 +++++++++++------- Mathlib/SetTheory/Cardinal/Aleph.lean | 4 - Mathlib/SetTheory/Cardinal/Cofinality.lean | 3 + 3 files changed, 135 insertions(+), 95 deletions(-) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Card.lean b/Mathlib/MeasureTheory/MeasurableSpace/Card.lean index f6d160c7eefc7..6598d8f936501 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Card.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Card.lean @@ -25,147 +25,188 @@ indeed generates this sigma-algebra. -/ -universe u +universe u v variable {α : Type u} -open Cardinal Set - --- Porting note: fix universe below, not here -local notation "ω₁" => (Ordinal.toType <| Cardinal.ord (aleph 1)) +open Cardinal Ordinal Set MeasureTheory namespace MeasurableSpace /-- Transfinite induction construction of the sigma-algebra generated by a set of sets `s`. At each step, we add all elements of `s`, the empty set, the complements of already constructed sets, and -countable unions of already constructed sets. We index this construction by an ordinal `< ω₁`, as -this will be enough to generate all sets in the sigma-algebra. +countable unions of already constructed sets. + +We index this construction by an arbitrary ordinal for simplicity, but by `ω₁` we will have +generated all the sets in the sigma-algebra. This construction is very similar to that of the Borel hierarchy. -/ -def generateMeasurableRec (s : Set (Set α)) (i : (ω₁ : Type u)) : Set (Set α) := - let S := ⋃ j : Iio i, generateMeasurableRec s (j.1) +def generateMeasurableRec (s : Set (Set α)) (i : Ordinal) : Set (Set α) := + let S := ⋃ j < i, generateMeasurableRec s j s ∪ {∅} ∪ compl '' S ∪ Set.range fun f : ℕ → S => ⋃ n, (f n).1 termination_by i -decreasing_by exact j.2 -theorem self_subset_generateMeasurableRec (s : Set (Set α)) (i : ω₁) : +theorem self_subset_generateMeasurableRec (s : Set (Set α)) (i : Ordinal) : s ⊆ generateMeasurableRec s i := by unfold generateMeasurableRec apply_rules [subset_union_of_subset_left] exact subset_rfl -theorem empty_mem_generateMeasurableRec (s : Set (Set α)) (i : ω₁) : +theorem empty_mem_generateMeasurableRec (s : Set (Set α)) (i : Ordinal) : ∅ ∈ generateMeasurableRec s i := by unfold generateMeasurableRec exact mem_union_left _ (mem_union_left _ (mem_union_right _ (mem_singleton ∅))) -theorem compl_mem_generateMeasurableRec {s : Set (Set α)} {i j : ω₁} (h : j < i) {t : Set α} +theorem compl_mem_generateMeasurableRec {s : Set (Set α)} {i j : Ordinal} (h : j < i) {t : Set α} (ht : t ∈ generateMeasurableRec s j) : tᶜ ∈ generateMeasurableRec s i := by unfold generateMeasurableRec - exact mem_union_left _ (mem_union_right _ ⟨t, mem_iUnion.2 ⟨⟨j, h⟩, ht⟩, rfl⟩) + exact mem_union_left _ (mem_union_right _ ⟨t, mem_iUnion₂.2 ⟨j, h, ht⟩, rfl⟩) -theorem iUnion_mem_generateMeasurableRec {s : Set (Set α)} {i : ω₁} {f : ℕ → Set α} +theorem iUnion_mem_generateMeasurableRec {s : Set (Set α)} {i : Ordinal} {f : ℕ → Set α} (hf : ∀ n, ∃ j < i, f n ∈ generateMeasurableRec s j) : - (⋃ n, f n) ∈ generateMeasurableRec s i := by + ⋃ n, f n ∈ generateMeasurableRec s i := by unfold generateMeasurableRec - exact mem_union_right _ ⟨fun n => ⟨f n, let ⟨j, hj, hf⟩ := hf n; mem_iUnion.2 ⟨⟨j, hj⟩, hf⟩⟩, rfl⟩ + exact mem_union_right _ ⟨fun n => ⟨f n, let ⟨j, hj, hf⟩ := hf n; mem_iUnion₂.2 ⟨j, hj, hf⟩⟩, rfl⟩ -theorem generateMeasurableRec_subset (s : Set (Set α)) {i j : ω₁} (h : i ≤ j) : - generateMeasurableRec s i ⊆ generateMeasurableRec s j := fun x hx => by - rcases eq_or_lt_of_le h with (rfl | h) +theorem generateMeasurableRec_mono (s : Set (Set α)) : Monotone (generateMeasurableRec s) := by + intro i j h x hx + rcases h.eq_or_lt with (rfl | h) · exact hx · convert iUnion_mem_generateMeasurableRec fun _ => ⟨i, h, hx⟩ exact (iUnion_const x).symm -/-- At each step of the inductive construction, the cardinality bound `≤ (max #s 2) ^ ℵ₀` holds. -/ -theorem cardinal_generateMeasurableRec_le (s : Set (Set α)) (i : ω₁) : +/-- An inductive principle for the elements of `generateMeasurableRec`. -/ +@[elab_as_elim] +theorem generateMeasurableRec_induction {s : Set (Set α)} {i : Ordinal} {t : Set α} + {p : Set α → Prop} (hs : ∀ t ∈ s, p t) (h0 : p ∅) + (hc : ∀ u, p u → (∃ j < i, u ∈ generateMeasurableRec s j) → p uᶜ) + (hn : ∀ f : ℕ → Set α, + (∀ n, p (f n) ∧ ∃ j < i, f n ∈ generateMeasurableRec s j) → p (⋃ n, f n)) : + t ∈ generateMeasurableRec s i → p t := by + suffices H : ∀ k ≤ i, ∀ t ∈ generateMeasurableRec s k, p t from H i le_rfl t + intro k + apply WellFoundedLT.induction k + intro k IH hk t + replace IH := fun j hj => IH j hj (hj.le.trans hk) + unfold generateMeasurableRec + rintro (((ht | rfl) | ht) | ⟨f, rfl⟩) + · exact hs t ht + · exact h0 + · simp_rw [mem_image, mem_iUnion₂] at ht + obtain ⟨u, ⟨⟨j, hj, hj'⟩, rfl⟩⟩ := ht + exact hc u (IH j hj u hj') ⟨j, hj.trans_le hk, hj'⟩ + · apply hn + intro n + obtain ⟨j, hj, hj'⟩ := mem_iUnion₂.1 (f n).2 + use IH j hj _ hj', j, hj.trans_le hk + +theorem generateMeasurableRec_omega1 (s : Set (Set α)) : + generateMeasurableRec s (ω₁ : Ordinal.{v}) = + ⋃ i < (ω₁ : Ordinal.{v}), generateMeasurableRec s i := by + apply (iUnion₂_subset fun i h => generateMeasurableRec_mono s h.le).antisymm' + intro t ht + rw [mem_iUnion₂] + refine generateMeasurableRec_induction ?_ ?_ ?_ ?_ ht + · intro t ht + exact ⟨0, omega_pos 1, self_subset_generateMeasurableRec s 0 ht⟩ + · exact ⟨0, omega_pos 1, empty_mem_generateMeasurableRec s 0⟩ + · rintro u - ⟨j, hj, hj'⟩ + exact ⟨_, (isLimit_omega 1).succ_lt hj, + compl_mem_generateMeasurableRec (Order.lt_succ j) hj'⟩ + · intro f H + choose I hI using fun n => (H n).1 + simp_rw [exists_prop] at hI + refine ⟨_, Ordinal.lsub_lt_ord_lift ?_ fun n => (hI n).1, + iUnion_mem_generateMeasurableRec fun n => ⟨_, Ordinal.lt_lsub I n, (hI n).2⟩⟩ + rw [mk_nat, lift_aleph0, isRegular_aleph_one.cof_omega_eq] + exact aleph0_lt_aleph_one + +theorem generateMeasurableRec_subset (s : Set (Set α)) (i : Ordinal) : + generateMeasurableRec s i ⊆ { t | GenerateMeasurable s t } := by + apply WellFoundedLT.induction i + exact fun i IH t ht => generateMeasurableRec_induction .basic .empty + (fun u _ ⟨j, hj, hj'⟩ => .compl _ (IH j hj hj')) (fun f H => .iUnion _ fun n => (H n).1) ht + +/-- `generateMeasurableRec s ω₁` generates precisely the smallest sigma-algebra containing `s`. -/ +theorem generateMeasurable_eq_rec (s : Set (Set α)) : + { t | GenerateMeasurable s t } = generateMeasurableRec s ω₁ := by + apply (generateMeasurableRec_subset s _).antisymm' + intro t ht + induction' ht with u hu u _ IH f _ IH + · exact self_subset_generateMeasurableRec s _ hu + · exact empty_mem_generateMeasurableRec s _ + · rw [generateMeasurableRec_omega1, mem_iUnion₂] at IH + obtain ⟨i, hi, hi'⟩ := IH + exact generateMeasurableRec_mono _ ((isLimit_omega 1).succ_lt hi).le + (compl_mem_generateMeasurableRec (Order.lt_succ i) hi') + · simp_rw [generateMeasurableRec_omega1, mem_iUnion₂, exists_prop] at IH + exact iUnion_mem_generateMeasurableRec IH + +/-- `generateMeasurableRec` is constant for ordinals `≥ ω₁`. -/ +theorem generateMeasurableRec_of_omega1_le (s : Set (Set α)) {i : Ordinal.{v}} (hi : ω₁ ≤ i) : + generateMeasurableRec s i = generateMeasurableRec s (ω₁ : Ordinal.{v}) := by + apply (generateMeasurableRec_mono s hi).antisymm' + rw [← generateMeasurable_eq_rec] + exact generateMeasurableRec_subset s i + +/-- At each step of the inductive construction, the cardinality bound `≤ #s ^ ℵ₀` holds. -/ +theorem cardinal_generateMeasurableRec_le (s : Set (Set α)) (i : Ordinal.{v}) : #(generateMeasurableRec s i) ≤ max #s 2 ^ ℵ₀ := by + suffices ∀ i ≤ ω₁, #(generateMeasurableRec s i) ≤ max #s 2 ^ ℵ₀ by + obtain hi | hi := le_or_lt i ω₁ + · exact this i hi + · rw [generateMeasurableRec_of_omega1_le s hi.le] + exact this _ le_rfl + intro i apply WellFoundedLT.induction i - intro i IH - have A := aleph0_le_aleph 1 - have B : aleph 1 ≤ max #s 2 ^ ℵ₀ := - aleph_one_le_continuum.trans (power_le_power_right (le_max_right _ _)) - have C : ℵ₀ ≤ max #s 2 ^ ℵ₀ := A.trans B - have J : #(⋃ j : Iio i, generateMeasurableRec s j.1) ≤ max #s 2 ^ ℵ₀ := by - refine (mk_iUnion_le _).trans ?_ - have D : ⨆ j : Iio i, #(generateMeasurableRec s j) ≤ _ := ciSup_le' fun ⟨j, hj⟩ => IH j hj - apply (mul_le_mul' ((mk_subtype_le _).trans (aleph 1).mk_ord_toType.le) D).trans - rw [mul_eq_max A C] - exact max_le B le_rfl + intro i IH hi + have A : 𝔠 ≤ max #s 2 ^ ℵ₀ := power_le_power_right (le_max_right _ _) + have B := aleph0_le_continuum.trans A + have C : #(⋃ j < i, generateMeasurableRec s j) ≤ max #s 2 ^ ℵ₀ := by + apply mk_iUnion_Ordinal_lift_le_of_le _ B _ + · intro j hj + exact IH j hj (hj.trans_le hi).le + · rw [lift_power, lift_aleph0] + rw [← Ordinal.lift_le.{u}, lift_omega, Ordinal.lift_one, ← ord_aleph] at hi + have H := card_le_of_le_ord hi + rw [← Ordinal.lift_card] at H + apply H.trans <| aleph_one_le_continuum.trans <| power_le_power_right _ + rw [lift_max, Cardinal.lift_ofNat] + exact le_max_right _ _ rw [generateMeasurableRec] - apply_rules [(mk_union_le _ _).trans, add_le_of_le C, mk_image_le.trans] - · exact (le_max_left _ _).trans (self_le_power _ one_lt_aleph0.le) + apply_rules [(mk_union_le _ _).trans, add_le_of_le (aleph_one_le_continuum.trans A), + mk_image_le.trans] + · exact (self_le_power _ one_le_aleph0).trans (power_le_power_right (le_max_left _ _)) · rw [mk_singleton] - exact one_lt_aleph0.le.trans C + exact one_lt_aleph0.le.trans B · apply mk_range_le.trans - simp only [mk_pi, prod_const, lift_uzero, mk_denumerable, lift_aleph0] - have := @power_le_power_right _ _ ℵ₀ J + simp only [mk_pi, prod_const, Cardinal.lift_uzero, mk_denumerable, lift_aleph0] + have := @power_le_power_right _ _ ℵ₀ C rwa [← power_mul, aleph0_mul_aleph0] at this -/-- `generateMeasurableRec s` generates precisely the smallest sigma-algebra containing `s`. -/ -theorem generateMeasurable_eq_rec (s : Set (Set α)) : - { t | GenerateMeasurable s t } = - ⋃ (i : (aleph 1).ord.toType), generateMeasurableRec s i := by - ext t; refine ⟨fun ht => ?_, fun ht => ?_⟩ - · inhabit ω₁ - induction ht with - | basic u hu => exact mem_iUnion.2 ⟨default, self_subset_generateMeasurableRec s _ hu⟩ - | empty => exact mem_iUnion.2 ⟨default, empty_mem_generateMeasurableRec s _⟩ - | compl _ _ IH => - rcases mem_iUnion.1 IH with ⟨i, hi⟩ - obtain ⟨j, hj⟩ := exists_gt i - exact mem_iUnion.2 ⟨j, compl_mem_generateMeasurableRec hj hi⟩ - | iUnion f _ IH => - have : ∀ n, ∃ i, f n ∈ generateMeasurableRec s i := fun n => by simpa using IH n - choose I hI using this - refine mem_iUnion.2 ⟨Ordinal.enum (α := ω₁) (· < ·) - ⟨Ordinal.lsub fun n => Ordinal.typein.{u} (α := ω₁) (· < ·) (I n), ?_⟩, - iUnion_mem_generateMeasurableRec fun n => ⟨I n, ?_, hI n⟩⟩ - · rw [Ordinal.type_toType] - refine Ordinal.lsub_lt_ord_lift ?_ fun i => Ordinal.typein_lt_self _ - rw [mk_denumerable, lift_aleph0, isRegular_aleph_one.cof_eq] - exact aleph0_lt_aleph_one - · rw [← Ordinal.typein_lt_typein (· < ·), Ordinal.typein_enum] - apply Ordinal.lt_lsub fun n : ℕ => _ - · rcases ht with ⟨t, ⟨i, rfl⟩, hx⟩ - revert t - apply WellFoundedLT.induction i - intro j H t ht - unfold generateMeasurableRec at ht - rcases ht with (((h | (rfl : t = ∅)) | ⟨u, ⟨-, ⟨⟨k, hk⟩, rfl⟩, hu⟩, rfl⟩) | ⟨f, rfl⟩) - · exact .basic t h - · exact .empty - · exact .compl u (H k hk u hu) - · refine .iUnion _ @fun n => ?_ - obtain ⟨-, ⟨⟨k, hk⟩, rfl⟩, hf⟩ := (f n).prop - exact H k hk _ hf - /-- If a sigma-algebra is generated by a set of sets `s`, then the sigma-algebra has cardinality at -most `(max #s 2) ^ ℵ₀`. -/ +most `max #s 2 ^ ℵ₀`. -/ theorem cardinal_generateMeasurable_le (s : Set (Set α)) : #{ t | GenerateMeasurable s t } ≤ max #s 2 ^ ℵ₀ := by - rw [generateMeasurable_eq_rec] - apply (mk_iUnion_le _).trans - rw [(aleph 1).mk_ord_toType] - refine le_trans (mul_le_mul' aleph_one_le_continuum - (ciSup_le' fun i => cardinal_generateMeasurableRec_le s i)) ?_ - refine (mul_le_max_of_aleph0_le_left aleph0_le_continuum).trans (max_le ?_ le_rfl) - exact power_le_power_right (le_max_right _ _) + rw [generateMeasurable_eq_rec.{u, 0}] + exact cardinal_generateMeasurableRec_le s _ /-- If a sigma-algebra is generated by a set of sets `s`, then the sigma -algebra has cardinality at most `(max #s 2) ^ ℵ₀`. -/ -theorem cardinalMeasurableSet_le (s : Set (Set α)) : +algebra has cardinality at most `max #s 2 ^ ℵ₀`. -/ +theorem cardinal_measurableSet_le (s : Set (Set α)) : #{ t | @MeasurableSet α (generateFrom s) t } ≤ max #s 2 ^ ℵ₀ := cardinal_generateMeasurable_le s +@[deprecated cardinal_measurableSet_le (since := "2024-08-30")] +alias cardinalMeasurableSet_le := cardinal_measurableSet_le + /-- If a sigma-algebra is generated by a set of sets `s` with cardinality at most the continuum, then the sigma algebra has the same cardinality bound. -/ theorem cardinal_generateMeasurable_le_continuum {s : Set (Set α)} (hs : #s ≤ 𝔠) : - #{ t | GenerateMeasurable s t } ≤ 𝔠 := - (cardinal_generateMeasurable_le s).trans - (by - rw [← continuum_power_aleph0] - exact mod_cast power_le_power_right (max_le hs (nat_lt_continuum 2).le)) + #{ t | GenerateMeasurable s t } ≤ 𝔠 := by + apply (cardinal_generateMeasurable_le s).trans + rw [← continuum_power_aleph0] + exact_mod_cast power_le_power_right (max_le hs (nat_lt_continuum 2).le) /-- If a sigma-algebra is generated by a set of sets `s` with cardinality at most the continuum, then the sigma algebra has the same cardinality bound. -/ diff --git a/Mathlib/SetTheory/Cardinal/Aleph.lean b/Mathlib/SetTheory/Cardinal/Aleph.lean index 5d5d56a75f0ee..f95620122c4e1 100644 --- a/Mathlib/SetTheory/Cardinal/Aleph.lean +++ b/Mathlib/SetTheory/Cardinal/Aleph.lean @@ -437,10 +437,6 @@ theorem isLimit_omega (o : Ordinal) : Ordinal.IsLimit (ω_ o) := by theorem ord_aleph_isLimit (o : Ordinal) : (ℵ_ o).ord.IsLimit := isLimit_ord <| aleph0_le_aleph _ --- TODO: get rid of this instance where it's used. -instance (o : Ordinal) : NoMaxOrder (ℵ_ o).ord.toType := - toType_noMax_of_succ_lt (isLimit_ord <| aleph0_le_aleph o).2 - @[simp] theorem range_aleph : range aleph = Set.Ici ℵ₀ := by ext c diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index d41d5f8e58c94..4024caa0c1f7a 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -899,6 +899,9 @@ theorem IsRegular.aleph0_le {c : Cardinal} (H : c.IsRegular) : ℵ₀ ≤ c := theorem IsRegular.cof_eq {c : Cardinal} (H : c.IsRegular) : c.ord.cof = c := (cof_ord_le c).antisymm H.2 +theorem IsRegular.cof_omega_eq {o : Ordinal} (H : (ℵ_ o).IsRegular) : (ω_ o).cof = ℵ_ o := by + rw [← ord_aleph, H.cof_eq] + theorem IsRegular.pos {c : Cardinal} (H : c.IsRegular) : 0 < c := aleph0_pos.trans_le H.1 From 812897a4585e90015527a14b09bd40d68c996be7 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Tue, 19 Nov 2024 00:27:42 +0000 Subject: [PATCH 093/829] feat(Manifold/IntegralCurve/Transform): pointwise notation and translation lemmas for subtraction (#19008) We restate the translation lemmas for integral curves using the `Pointwise` API and add translation lemmas for subtraction for convenience. This is just split out from #9013. The `Pointwise` API allows us to use lemmas like `Metric.vadd_ball` and more convenient rewriting of the `dt` term (rather than rewriting it inside the set builder notation). --- .../Manifold/IntegralCurve/Transform.lean | 27 ++++++++++++++----- 1 file changed, 20 insertions(+), 7 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve/Transform.lean b/Mathlib/Geometry/Manifold/IntegralCurve/Transform.lean index c60216be7a1df..f25fae3a60a20 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve/Transform.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve/Transform.lean @@ -20,7 +20,7 @@ curve. integral curve, vector field -/ -open Function Set +open Function Set Pointwise variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] @@ -33,21 +33,27 @@ variable section Translation lemma IsIntegralCurveOn.comp_add (hγ : IsIntegralCurveOn γ v s) (dt : ℝ) : - IsIntegralCurveOn (γ ∘ (· + dt)) v { t | t + dt ∈ s } := by + IsIntegralCurveOn (γ ∘ (· + dt)) v (-dt +ᵥ s) := by intros t ht rw [comp_apply, ← ContinuousLinearMap.comp_id (ContinuousLinearMap.smulRight 1 (v (γ (t + dt))))] + rw [mem_vadd_set_iff_neg_vadd_mem, neg_neg, vadd_eq_add, add_comm, + ← mem_setOf (p := fun t ↦ t + dt ∈ s)] at ht apply HasMFDerivAt.comp t (hγ (t + dt) ht) refine ⟨(continuous_add_right _).continuousAt, ?_⟩ simp only [mfld_simps, hasFDerivWithinAt_univ] exact HasFDerivAt.add_const (hasFDerivAt_id _) _ lemma isIntegralCurveOn_comp_add {dt : ℝ} : - IsIntegralCurveOn γ v s ↔ IsIntegralCurveOn (γ ∘ (· + dt)) v { t | t + dt ∈ s } := by + IsIntegralCurveOn γ v s ↔ IsIntegralCurveOn (γ ∘ (· + dt)) v (-dt +ᵥ s) := by refine ⟨fun hγ ↦ hγ.comp_add _, fun hγ ↦ ?_⟩ convert hγ.comp_add (-dt) · ext t simp only [Function.comp_apply, neg_add_cancel_right] - · simp only [mem_setOf_eq, neg_add_cancel_right, setOf_mem_eq] + · simp only [neg_neg, vadd_neg_vadd] + +lemma isIntegralCurveOn_comp_sub {dt : ℝ} : + IsIntegralCurveOn γ v s ↔ IsIntegralCurveOn (γ ∘ (· - dt)) v (dt +ᵥ s) := by + simpa using isIntegralCurveOn_comp_add (dt := -dt) lemma IsIntegralCurveAt.comp_add (hγ : IsIntegralCurveAt γ v t₀) (dt : ℝ) : IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by @@ -55,8 +61,7 @@ lemma IsIntegralCurveAt.comp_add (hγ : IsIntegralCurveAt γ v t₀) (dt : ℝ) obtain ⟨ε, hε, h⟩ := hγ refine ⟨ε, hε, ?_⟩ convert h.comp_add dt - ext t - rw [mem_setOf_eq, Metric.mem_ball, Metric.mem_ball, dist_sub_eq_dist_add_right] + rw [Metric.vadd_ball, vadd_eq_add, neg_add_eq_sub] lemma isIntegralCurveAt_comp_add {dt : ℝ} : IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by @@ -66,10 +71,14 @@ lemma isIntegralCurveAt_comp_add {dt : ℝ} : simp only [Function.comp_apply, neg_add_cancel_right] · simp only [sub_neg_eq_add, sub_add_cancel] +lemma isIntegralCurveAt_comp_sub {dt : ℝ} : + IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· - dt)) v (t₀ + dt) := by + simpa using isIntegralCurveAt_comp_add (dt := -dt) + lemma IsIntegralCurve.comp_add (hγ : IsIntegralCurve γ v) (dt : ℝ) : IsIntegralCurve (γ ∘ (· + dt)) v := by rw [isIntegralCurve_iff_isIntegralCurveOn] at * - exact hγ.comp_add _ + simpa using hγ.comp_add dt lemma isIntegralCurve_comp_add {dt : ℝ} : IsIntegralCurve γ v ↔ IsIntegralCurve (γ ∘ (· + dt)) v := by @@ -78,6 +87,10 @@ lemma isIntegralCurve_comp_add {dt : ℝ} : ext t simp only [Function.comp_apply, neg_add_cancel_right] +lemma isIntegralCurve_comp_sub {dt : ℝ} : + IsIntegralCurve γ v ↔ IsIntegralCurve (γ ∘ (· - dt)) v := by + simpa using isIntegralCurve_comp_add (dt := -dt) + end Translation /-! ### Scaling lemmas -/ From 47f0d27ac082d7d45d4efcca5442d6adcbcc0e2b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 19 Nov 2024 02:51:22 +0000 Subject: [PATCH 094/829] chore(SetTheory/Cardinal/Aleph): golf `aleph` theorems (#18204) --- Mathlib/SetTheory/Cardinal/Aleph.lean | 47 ++++++++++++++------------- 1 file changed, 24 insertions(+), 23 deletions(-) diff --git a/Mathlib/SetTheory/Cardinal/Aleph.lean b/Mathlib/SetTheory/Cardinal/Aleph.lean index f95620122c4e1..3970c6fc6cb57 100644 --- a/Mathlib/SetTheory/Cardinal/Aleph.lean +++ b/Mathlib/SetTheory/Cardinal/Aleph.lean @@ -207,6 +207,9 @@ theorem omega_le_omega {o₁ o₂ : Ordinal} : ω_ o₁ ≤ ω_ o₂ ↔ o₁ theorem omega_max (o₁ o₂ : Ordinal) : ω_ (max o₁ o₂) = max (ω_ o₁) (ω_ o₂) := omega.monotone.map_max +theorem preOmega_le_omega (o : Ordinal) : preOmega o ≤ ω_ o := + preOmega_le_preOmega.2 (Ordinal.le_add_left _ _) + theorem isInitial_omega (o : Ordinal) : IsInitial (omega o) := isInitial_preOmega _ @@ -298,13 +301,21 @@ theorem preAleph_succ (o : Ordinal) : preAleph (succ o) = succ (preAleph o) := preAleph.map_succ o @[simp] -theorem preAleph_nat : ∀ n : ℕ, preAleph n = n - | 0 => preAleph_zero - | n + 1 => show preAleph (succ n) = n.succ by rw [preAleph_succ, preAleph_nat n, nat_succ] +theorem preAleph_nat (n : ℕ) : preAleph n = n := by + rw [← card_preOmega, preOmega_natCast, card_nat] + +@[simp] +theorem preAleph_omega0 : preAleph ω = ℵ₀ := by + rw [← card_preOmega, preOmega_omega0, card_omega0] +@[simp] theorem preAleph_pos {o : Ordinal} : 0 < preAleph o ↔ 0 < o := by rw [← preAleph_zero, preAleph_lt_preAleph] +@[simp] +theorem aleph0_le_preAleph {o : Ordinal} : ℵ₀ ≤ preAleph o ↔ ω ≤ o := by + rw [← preAleph_omega0, preAleph_le_preAleph] + @[simp] theorem lift_preAleph (o : Ordinal.{u}) : lift.{v} (preAleph o) = preAleph (Ordinal.lift.{v} o) := (preAleph.toInitialSeg.trans liftInitialSeg).eq @@ -328,16 +339,6 @@ theorem preAleph_limit {o : Ordinal} (ho : o.IsLimit) : preAleph o = ⨆ a : Iio rw [preAleph_le_of_isLimit ho] exact fun a ha => le_ciSup (bddAbove_of_small _) (⟨a, ha⟩ : Iio o) -@[simp] -theorem preAleph_omega0 : preAleph ω = ℵ₀ := - eq_of_forall_ge_iff fun c => by - simp only [preAleph_le_of_isLimit isLimit_omega0, lt_omega0, exists_imp, aleph0_le] - exact forall_swap.trans (forall_congr' fun n => by simp only [forall_eq, preAleph_nat]) - -@[simp] -theorem aleph0_le_preAleph {o : Ordinal} : ℵ₀ ≤ preAleph o ↔ ω ≤ o := by - rw [← preAleph_omega0, preAleph_le_preAleph] - /-- The `aleph` function gives the infinite cardinals listed by their ordinal index. `aleph 0 = ℵ₀`, `aleph 1 = succ ℵ₀` is the first uncountable cardinal, and so on. @@ -381,6 +382,9 @@ theorem aleph_max (o₁ o₂ : Ordinal) : ℵ_ (max o₁ o₂) = max (ℵ_ o₁) theorem max_aleph_eq (o₁ o₂ : Ordinal) : max (ℵ_ o₁) (ℵ_ o₂) = ℵ_ (max o₁ o₂) := (aleph_max o₁ o₂).symm +theorem preAleph_le_aleph (o : Ordinal) : preAleph o ≤ ℵ_ o := + preAleph_le_preAleph.2 (Ordinal.le_add_left _ _) + @[simp] theorem aleph_succ (o : Ordinal) : ℵ_ (succ o) = succ (ℵ_ o) := by rw [aleph_eq_preAleph, add_succ, preAleph_succ, aleph_eq_preAleph] @@ -399,16 +403,13 @@ theorem _root_.Ordinal.lift_omega (o : Ordinal.{u}) : simp [omega_eq_preOmega] theorem aleph_limit {o : Ordinal} (ho : o.IsLimit) : ℵ_ o = ⨆ a : Iio o, ℵ_ a := by - apply le_antisymm _ (ciSup_le' _) - · rw [aleph_eq_preAleph, preAleph_limit (ho.add _)] - refine ciSup_mono' (bddAbove_of_small _) ?_ - rintro ⟨i, hi⟩ - cases' lt_or_le i ω with h h - · rcases lt_omega0.1 h with ⟨n, rfl⟩ - use ⟨0, ho.pos⟩ - simpa using (nat_lt_aleph0 n).le - · exact ⟨⟨_, (sub_lt_of_le h).2 hi⟩, preAleph_le_preAleph.2 (le_add_sub _ _)⟩ - · exact fun i => aleph_le_aleph.2 i.2.le + rw [aleph_eq_preAleph, preAleph_limit (isLimit_add ω ho)] + apply le_antisymm <;> + apply ciSup_mono' (bddAbove_of_small _) <;> + intro i + · refine ⟨⟨_, sub_lt_of_lt_add i.2 ho.pos⟩, ?_⟩ + simpa [aleph_eq_preAleph] using le_add_sub _ _ + · exact ⟨⟨_, add_lt_add_left i.2 ω⟩, le_rfl⟩ theorem aleph0_le_aleph (o : Ordinal) : ℵ₀ ≤ ℵ_ o := by rw [aleph_eq_preAleph, aleph0_le_preAleph] From b2244f6fc5832d2e74f3990726e448489cb773bd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 19 Nov 2024 04:07:31 +0000 Subject: [PATCH 095/829] =?UTF-8?q?chore(SetTheory/Ordinal/Basic):=20`Ordi?= =?UTF-8?q?nal.NeZero.One`=20=E2=86=92=20`Ordinal.instNeZeroOne`=20(#19064?= =?UTF-8?q?)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/SetTheory/Nimber/Basic.lean | 14 ++++++------ Mathlib/SetTheory/Ordinal/Basic.lean | 14 ++++++------ Mathlib/SetTheory/Ordinal/NaturalOps.lean | 26 ++++++++++++----------- 3 files changed, 28 insertions(+), 26 deletions(-) diff --git a/Mathlib/SetTheory/Nimber/Basic.lean b/Mathlib/SetTheory/Nimber/Basic.lean index 28ff884f9c685..9f33020c6c6b5 100644 --- a/Mathlib/SetTheory/Nimber/Basic.lean +++ b/Mathlib/SetTheory/Nimber/Basic.lean @@ -44,16 +44,16 @@ noncomputable section /-! ### Basic casts between `Ordinal` and `Nimber` -/ -/-- A type synonym for ordinals with natural addition and multiplication. -/ +/-- A type synonym for ordinals with nimber addition and multiplication. -/ def Nimber : Type _ := Ordinal deriving Zero, Inhabited, One, Nontrivial, WellFoundedRelation -instance Nimber.linearOrder : LinearOrder Nimber := {Ordinal.linearOrder with} -instance Nimber.succOrder : SuccOrder Nimber := {Ordinal.instSuccOrder with} -instance Nimber.orderBot : OrderBot Nimber := {Ordinal.orderBot with} -instance Nimber.noMaxOrder : NoMaxOrder Nimber := {Ordinal.noMaxOrder with} -instance Nimber.zeroLEOneClass : ZeroLEOneClass Nimber := {Ordinal.zeroLEOneClass with} -instance Nimber.instNeZeroOne : NeZero (1 : Nimber) := Ordinal.NeZero.one +instance Nimber.instLinearOrder : LinearOrder Nimber := Ordinal.instLinearOrder +instance Nimber.instSuccOrder : SuccOrder Nimber := Ordinal.instSuccOrder +instance Nimber.instOrderBot : OrderBot Nimber := Ordinal.instOrderBot +instance Nimber.instNoMaxOrder : NoMaxOrder Nimber := Ordinal.instNoMaxOrder +instance Nimber.instZeroLEOneClass : ZeroLEOneClass Nimber := Ordinal.instZeroLEOneClass +instance Nimber.instNeZeroOne : NeZero (1 : Nimber) := Ordinal.instNeZeroOne /-- The identity function between `Ordinal` and `Nimber`. -/ @[match_pattern] diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index 65e5733c47823..5a31bfe919ee9 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -287,7 +287,7 @@ instance partialOrder : PartialOrder Ordinal where Quotient.inductionOn₂ a b fun _ _ ⟨h₁⟩ ⟨h₂⟩ => Quot.sound ⟨InitialSeg.antisymm h₁ h₂⟩ -instance linearOrder : LinearOrder Ordinal := +instance : LinearOrder Ordinal := {inferInstanceAs (PartialOrder Ordinal) with le_total := fun a b => Quotient.inductionOn₂ a b fun ⟨_, r, _⟩ ⟨_, s, _⟩ => (InitialSeg.total r s).recOn (fun f => Or.inl ⟨f⟩) fun f => Or.inr ⟨f⟩ @@ -309,7 +309,7 @@ theorem _root_.PrincipalSeg.ordinal_type_lt {α β} {r : α → α → Prop} {s protected theorem zero_le (o : Ordinal) : 0 ≤ o := inductionOn o fun _ r _ => (InitialSeg.ofIsEmpty _ r).ordinal_type_le -instance orderBot : OrderBot Ordinal where +instance : OrderBot Ordinal where bot := 0 bot_le := Ordinal.zero_le @@ -330,10 +330,10 @@ protected theorem not_lt_zero (o : Ordinal) : ¬o < 0 := theorem eq_zero_or_pos : ∀ a : Ordinal, a = 0 ∨ 0 < a := eq_bot_or_bot_lt -instance zeroLEOneClass : ZeroLEOneClass Ordinal := +instance : ZeroLEOneClass Ordinal := ⟨Ordinal.zero_le _⟩ -instance NeZero.one : NeZero (1 : Ordinal) := +instance instNeZeroOne : NeZero (1 : Ordinal) := ⟨Ordinal.one_ne_zero⟩ theorem type_le_iff {α β} {r : α → α → Prop} {s : β → β → Prop} [IsWellOrder α r] @@ -875,13 +875,13 @@ private theorem succ_le_iff' {a b : Ordinal} : a + 1 ≤ b ↔ a < b := by · apply (RelEmbedding.ofMonotone (Sum.recOn · f fun _ ↦ f.top) ?_).ordinal_type_le simpa [f.map_rel_iff] using f.lt_top -instance noMaxOrder : NoMaxOrder Ordinal := +instance : NoMaxOrder Ordinal := ⟨fun _ => ⟨_, succ_le_iff'.1 le_rfl⟩⟩ -instance instSuccOrder : SuccOrder Ordinal.{u} := +instance : SuccOrder Ordinal.{u} := SuccOrder.ofSuccLeIff (fun o => o + 1) succ_le_iff' -instance instSuccAddOrder : SuccAddOrder Ordinal := ⟨fun _ => rfl⟩ +instance : SuccAddOrder Ordinal := ⟨fun _ => rfl⟩ @[simp] theorem add_one_eq_succ (o : Ordinal) : o + 1 = succ o := diff --git a/Mathlib/SetTheory/Ordinal/NaturalOps.lean b/Mathlib/SetTheory/Ordinal/NaturalOps.lean index c7ac9850362c8..7463af175343e 100644 --- a/Mathlib/SetTheory/Ordinal/NaturalOps.lean +++ b/Mathlib/SetTheory/Ordinal/NaturalOps.lean @@ -53,10 +53,12 @@ def NatOrdinal : Type _ := -- Porting note: used to derive LinearOrder & SuccOrder but need to manually define Ordinal deriving Zero, Inhabited, One, WellFoundedRelation -instance NatOrdinal.linearOrder : LinearOrder NatOrdinal := {Ordinal.linearOrder with} -instance NatOrdinal.instSuccOrder : SuccOrder NatOrdinal := {Ordinal.instSuccOrder with} -instance NatOrdinal.orderBot : OrderBot NatOrdinal := {Ordinal.orderBot with} -instance NatOrdinal.noMaxOrder : NoMaxOrder NatOrdinal := {Ordinal.noMaxOrder with} +instance NatOrdinal.instLinearOrder : LinearOrder NatOrdinal := Ordinal.instLinearOrder +instance NatOrdinal.instSuccOrder : SuccOrder NatOrdinal := Ordinal.instSuccOrder +instance NatOrdinal.instOrderBot : OrderBot NatOrdinal := Ordinal.instOrderBot +instance NatOrdinal.instNoMaxOrder : NoMaxOrder NatOrdinal := Ordinal.instNoMaxOrder +instance NatOrdinal.instZeroLEOneClass : ZeroLEOneClass NatOrdinal := Ordinal.instZeroLEOneClass +instance NatOrdinal.instNeZeroOne : NeZero (1 : NatOrdinal) := Ordinal.instNeZeroOne /-- The identity function between `Ordinal` and `NatOrdinal`. -/ @[match_pattern] @@ -316,19 +318,19 @@ open Ordinal NaturalOps instance : Add NatOrdinal := ⟨nadd⟩ instance : SuccAddOrder NatOrdinal := ⟨fun x => (nadd_one x).symm⟩ -instance addLeftStrictMono : AddLeftStrictMono NatOrdinal.{u} := +instance : AddLeftStrictMono NatOrdinal.{u} := ⟨fun a _ _ h => nadd_lt_nadd_left h a⟩ -instance addLeftMono : AddLeftMono NatOrdinal.{u} := +instance : AddLeftMono NatOrdinal.{u} := ⟨fun a _ _ h => nadd_le_nadd_left h a⟩ -instance addLeftReflectLE : AddLeftReflectLE NatOrdinal.{u} := +instance : AddLeftReflectLE NatOrdinal.{u} := ⟨fun a b c h => by by_contra! h' exact h.not_lt (add_lt_add_left h' a)⟩ -instance orderedCancelAddCommMonoid : OrderedCancelAddCommMonoid NatOrdinal := - { NatOrdinal.linearOrder with +instance : OrderedCancelAddCommMonoid NatOrdinal := + { NatOrdinal.instLinearOrder with add := (· + ·) add_assoc := nadd_assoc add_le_add_left := fun _ _ => add_le_add_left @@ -339,7 +341,7 @@ instance orderedCancelAddCommMonoid : OrderedCancelAddCommMonoid NatOrdinal := add_comm := nadd_comm nsmul := nsmulRec } -instance addMonoidWithOne : AddMonoidWithOne NatOrdinal := +instance : AddMonoidWithOne NatOrdinal := AddMonoidWithOne.unary @[deprecated Order.succ_eq_add_one (since := "2024-09-04")] @@ -674,8 +676,8 @@ instance : Mul NatOrdinal := -- Porting note: had to add universe annotations to ensure that the -- two sources lived in the same universe. instance : OrderedCommSemiring NatOrdinal.{u} := - { NatOrdinal.orderedCancelAddCommMonoid.{u}, - NatOrdinal.linearOrder.{u} with + { NatOrdinal.instOrderedCancelAddCommMonoid.{u}, + NatOrdinal.instLinearOrder.{u} with mul := (· * ·) left_distrib := nmul_nadd right_distrib := nadd_nmul From 955e8f97a6372ceeeb97f4acc87f71ae1fea7d85 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 19 Nov 2024 05:23:28 +0000 Subject: [PATCH 096/829] chore: fix defeq abuse between `List.get_mem` and `List.getElem_mem` (#19224) In all these places the goal state is about `getElem` not `List.get`. This will reduce the fallout of leanprover/lean4#6095. --- Mathlib/Data/List/EditDistance/Bounds.lean | 2 +- Mathlib/Data/List/MinMax.lean | 2 +- Mathlib/Data/List/Nodup.lean | 2 +- Mathlib/Data/List/NodupEquivFin.lean | 2 +- Mathlib/Data/List/Permutation.lean | 4 ++-- 5 files changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Data/List/EditDistance/Bounds.lean b/Mathlib/Data/List/EditDistance/Bounds.lean index 7a20e9e681c9a..df199af6ccfa1 100644 --- a/Mathlib/Data/List/EditDistance/Bounds.lean +++ b/Mathlib/Data/List/EditDistance/Bounds.lean @@ -79,7 +79,7 @@ theorem le_suffixLevenshtein_append_minimum (xs : List α) (ys₁ ys₂) : theorem suffixLevenshtein_minimum_le_levenshtein_append (xs ys₁ ys₂) : (suffixLevenshtein C xs ys₂).1.minimum ≤ levenshtein C xs (ys₁ ++ ys₂) := by cases ys₁ with - | nil => exact List.minimum_le_of_mem' (List.get_mem _ _ _) + | nil => exact List.minimum_le_of_mem' (List.getElem_mem _) | cons y ys₁ => exact (le_suffixLevenshtein_append_minimum _ _ _).trans (suffixLevenshtein_minimum_le_levenshtein_cons _ _ _) diff --git a/Mathlib/Data/List/MinMax.lean b/Mathlib/Data/List/MinMax.lean index 6b18d25cf4107..0237438fa2eda 100644 --- a/Mathlib/Data/List/MinMax.lean +++ b/Mathlib/Data/List/MinMax.lean @@ -427,7 +427,7 @@ theorem minimum_of_length_pos_le_of_mem (h : a ∈ l) (w : 0 < l.length) : theorem getElem_le_maximum_of_length_pos {i : ℕ} (w : i < l.length) (h := (Nat.zero_lt_of_lt w)) : l[i] ≤ l.maximum_of_length_pos h := by apply le_maximum_of_length_pos_of_mem - exact get_mem l i w + exact getElem_mem _ theorem minimum_of_length_pos_le_getElem {i : ℕ} (w : i < l.length) (h := (Nat.zero_lt_of_lt w)) : l.minimum_of_length_pos h ≤ l[i] := diff --git a/Mathlib/Data/List/Nodup.lean b/Mathlib/Data/List/Nodup.lean index f6b34a9b86b77..1368f8f1890a3 100644 --- a/Mathlib/Data/List/Nodup.lean +++ b/Mathlib/Data/List/Nodup.lean @@ -122,7 +122,7 @@ theorem not_nodup_of_get_eq_of_ne (xs : List α) (n m : Fin xs.length) theorem indexOf_getElem [DecidableEq α] {l : List α} (H : Nodup l) (i : Nat) (h : i < l.length) : indexOf l[i] l = i := - suffices (⟨indexOf l[i] l, indexOf_lt_length.2 (get_mem _ _ _)⟩ : Fin l.length) = ⟨i, h⟩ + suffices (⟨indexOf l[i] l, indexOf_lt_length.2 (getElem_mem _)⟩ : Fin l.length) = ⟨i, h⟩ from Fin.val_eq_of_eq this nodup_iff_injective_get.1 H (by simp) diff --git a/Mathlib/Data/List/NodupEquivFin.lean b/Mathlib/Data/List/NodupEquivFin.lean index 0dcf11280caae..65a48a57e4574 100644 --- a/Mathlib/Data/List/NodupEquivFin.lean +++ b/Mathlib/Data/List/NodupEquivFin.lean @@ -127,7 +127,7 @@ theorem sublist_of_orderEmbedding_get?_eq {l l' : List α} (f : ℕ ↪o ℕ) 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 _)] - apply List.get_mem + exact List.getElem_mem _ /-- A `l : List α` is `Sublist l l'` for `l' : List α` iff there is `f`, an order-preserving embedding of `ℕ` into `ℕ` such that diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index d90a8bbe313e9..2ce7b4ce525f3 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -501,13 +501,13 @@ theorem nodup_permutations (s : List α) (hs : Nodup s) : Nodup s.permutations : rcases lt_trichotomy n m with (ht | ht | ht) · suffices x ∈ bs by exact h x (hb.subset this) rfl rw [← hx', getElem_insertIdx_of_lt _ _ _ _ ht (ht.trans_le hm)] - exact get_mem _ _ _ + exact getElem_mem _ · simp only [ht] at hm' hn' rw [← hm'] at hn' exact H (insertIdx_injective _ _ hn') · suffices x ∈ as by exact h x (ha.subset this) rfl rw [← hx, getElem_insertIdx_of_lt _ _ _ _ ht (ht.trans_le hn)] - exact get_mem _ _ _ + exact getElem_mem _ lemma permutations_take_two (x y : α) (s : List α) : (x :: y :: s).permutations.take 2 = [x :: y :: s, y :: x :: s] := by From 9274ddbe2cedc8b212a9e0d0dc92094a830c7abc Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Tue, 19 Nov 2024 07:30:29 +0000 Subject: [PATCH 097/829] feat(Tactic/StacksAttribute): improve comment shown in docstring for stack tags (#19235) Previously it was ```text Stacks Tag 09HK Part 1, finSepDegree variant ``` Now it should be ```text Stacks Tag 09HK (Part 1, finSepDegree variant) ``` which is more compact. --- Mathlib/Tactic/StacksAttribute.lean | 3 ++- MathlibTest/StacksAttribute.lean | 9 +++++++++ 2 files changed, 11 insertions(+), 1 deletion(-) diff --git a/Mathlib/Tactic/StacksAttribute.lean b/Mathlib/Tactic/StacksAttribute.lean index c8f7d829151e7..c7f69363dca6e 100644 --- a/Mathlib/Tactic/StacksAttribute.lean +++ b/Mathlib/Tactic/StacksAttribute.lean @@ -153,7 +153,8 @@ initialize Lean.registerBuiltinAttribute { | _ => throwUnsupportedSyntax let tagStr ← tag.getStacksTag let comment := (comment.map (·.getString)).getD "" - let newDoc := [oldDoc, s!"[{SorK} Tag {tagStr}]({url}/{tagStr})", comment] + let commentInDoc := if comment = "" then "" else s!" ({comment})" + let newDoc := [oldDoc, s!"[{SorK} Tag {tagStr}]({url}/{tagStr}){commentInDoc}"] addDocString decl <| "\n\n".intercalate (newDoc.filter (· != "")) addTagEntry decl database tagStr <| comment } diff --git a/MathlibTest/StacksAttribute.lean b/MathlibTest/StacksAttribute.lean index 92fe5efb40fcd..3b70bdb597610 100644 --- a/MathlibTest/StacksAttribute.lean +++ b/MathlibTest/StacksAttribute.lean @@ -11,6 +11,15 @@ theorem tagged : True := .intro end X +/-- +info: some ([Stacks Tag A04Q](https://stacks.math.columbia.edu/tag/A04Q) (A comment) + +[Kerodon Tag B15R](https://kerodon.net/tag/B15R) (Also a comment)) +-/ +#guard_msgs in +run_cmd + Lean.logInfo m!"{← Lean.findDocString? (← Lean.getEnv) `X.tagged}" + #guard_msgs in @[stacks 0BR2, kerodon 0X12] example : True := .intro From 3a1c2a7c444843bcfdea46b41c78b932217f5612 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Tue, 19 Nov 2024 08:16:25 +0000 Subject: [PATCH 098/829] feat(LinearAlgebra/Matrix): Permanent (#18936) Co-authored-by: Moritz Firsching --- Mathlib.lean | 1 + Mathlib/LinearAlgebra/Matrix/Permanent.lean | 86 +++++++++++++++++++++ 2 files changed, 87 insertions(+) create mode 100644 Mathlib/LinearAlgebra/Matrix/Permanent.lean diff --git a/Mathlib.lean b/Mathlib.lean index e9bea276ba9f6..7566572fa39e6 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3327,6 +3327,7 @@ import Mathlib.LinearAlgebra.Matrix.MvPolynomial import Mathlib.LinearAlgebra.Matrix.Nondegenerate import Mathlib.LinearAlgebra.Matrix.NonsingularInverse import Mathlib.LinearAlgebra.Matrix.Orthogonal +import Mathlib.LinearAlgebra.Matrix.Permanent import Mathlib.LinearAlgebra.Matrix.Permutation import Mathlib.LinearAlgebra.Matrix.Polynomial import Mathlib.LinearAlgebra.Matrix.PosDef diff --git a/Mathlib/LinearAlgebra/Matrix/Permanent.lean b/Mathlib/LinearAlgebra/Matrix/Permanent.lean new file mode 100644 index 0000000000000..fa1620053c90a --- /dev/null +++ b/Mathlib/LinearAlgebra/Matrix/Permanent.lean @@ -0,0 +1,86 @@ +/- +Copyright (c) 2024 Moritz Firsching. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Moritz Firsching +-/ +import Mathlib.Algebra.BigOperators.GroupWithZero.Finset +import Mathlib.Data.Fintype.Perm +import Mathlib.Data.Matrix.Diagonal +/-! +# Permanent of a matrix + +This file defines the permanent of a matrix, `Matrix.permanent`, and some of its properties. + +## Main definitions + +* `Matrix.permanent`: the permanent of a square matrix, as a sum over permutations + +-/ + +open Equiv Fintype Finset + +namespace Matrix + +variable {n : Type*} [DecidableEq n] [Fintype n] +variable {R : Type*} [CommSemiring R] + +/-- The permanent of a square matrix defined as a sum over all permutations. This is analogous to +the determinant but without alternating signs. -/ +def permanent (M : Matrix n n R) : R := ∑ σ : Perm n, ∏ i, M (σ i) i + +@[simp] +theorem permanent_diagonal {d : n → R} : permanent (diagonal d) = ∏ i, d i := by + refine (sum_eq_single 1 (fun σ _ hσ ↦ ?_) (fun h ↦ (h <| mem_univ _).elim)).trans ?_ + · match not_forall.mp (mt Equiv.ext hσ) with + | ⟨x, hx⟩ => exact Finset.prod_eq_zero (mem_univ x) (if_neg hx) + · simp only [Perm.one_apply, diagonal_apply_eq] + +@[simp] +theorem permanent_zero [Nonempty n] : permanent (0 : Matrix n n R) = 0 := by simp [permanent] + +@[simp] +theorem permanent_one : permanent (1 : Matrix n n R) = 1 := by + rw [← diagonal_one]; simp [-diagonal_one] + +theorem permanent_isEmpty [IsEmpty n] {A : Matrix n n R} : permanent A = 1 := by simp [permanent] + +theorem permanent_eq_one_of_card_eq_zero {A : Matrix n n R} (h : card n = 0) : permanent A = 1 := + haveI : IsEmpty n := card_eq_zero_iff.mp h + permanent_isEmpty + +/-- If `n` has only one element, the permanent of an `n` by `n` matrix is just that element. +Although `Unique` implies `DecidableEq` and `Fintype`, the instances might +not be syntactically equal. Thus, we need to fill in the args explicitly. -/ +@[simp] +theorem permanent_unique {n : Type*} [Unique n] [DecidableEq n] [Fintype n] (A : Matrix n n R) : + permanent A = A default default := by simp [permanent, univ_unique] + +theorem permanent_eq_elem_of_subsingleton [Subsingleton n] (A : Matrix n n R) (k : n) : + permanent A = A k k := by + have := uniqueOfSubsingleton k + convert permanent_unique A + +theorem permanent_eq_elem_of_card_eq_one {A : Matrix n n R} (h : card n = 1) (k : n) : + permanent A = A k k := + haveI : Subsingleton n := card_le_one_iff_subsingleton.mp h.le + permanent_eq_elem_of_subsingleton _ _ + +/-- Transposing a matrix preserves the permanent. -/ +@[simp] +theorem permanent_transpose (M : Matrix n n R) : Mᵀ.permanent = M.permanent := by + refine sum_bijective _ inv_involutive.bijective _ _ ?_ + intro σ + apply Fintype.prod_equiv σ + simp + +/-- Permuting the columns does not change the permanent. -/ +theorem permanent_permute_cols (σ : Perm n) (M : Matrix n n R) : + (M.submatrix σ id).permanent = M.permanent := + (Group.mulLeft_bijective σ).sum_comp fun τ ↦ ∏ i : n, M (τ i) i + +/-- Permuting the rows does not change the permanent. -/ +theorem permanent_permute_rows (σ : Perm n) (M : Matrix n n R) : + (M.submatrix id σ).permanent = M.permanent := by + rw [← permanent_transpose, transpose_submatrix, permanent_permute_cols, permanent_transpose] + +end Matrix From 809872da938e18d9d335b6c61e41f9c3cd4fac25 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Tue, 19 Nov 2024 08:25:54 +0000 Subject: [PATCH 099/829] refactor(CategoryTheory): make PreservesLimit and ReflectsLimit props (#19206) All changes are straightforward. Some universe parameters had to be made explicit, but without changing the level of generality of the definitions/statements. --- .../Algebra/Category/AlgebraCat/Limits.lean | 24 +- Mathlib/Algebra/Category/Grp/AB5.lean | 4 +- .../Category/Grp/FilteredColimits.lean | 18 +- Mathlib/Algebra/Category/Grp/Limits.lean | 60 +- .../Algebra/Category/ModuleCat/Abelian.lean | 16 +- .../Category/ModuleCat/ChangeOfRings.lean | 11 +- .../Algebra/Category/ModuleCat/Colimits.lean | 4 +- .../Category/ModuleCat/FilteredColimits.lean | 12 +- .../Algebra/Category/ModuleCat/Limits.lean | 34 +- .../Category/ModuleCat/Presheaf/Colimits.lean | 16 +- .../Category/ModuleCat/Presheaf/Limits.lean | 16 +- .../ModuleCat/Presheaf/Sheafification.lean | 4 +- .../Category/ModuleCat/Sheaf/Limits.lean | 4 +- .../Category/MonCat/FilteredColimits.lean | 12 +- Mathlib/Algebra/Category/MonCat/Limits.lean | 30 +- .../Category/Ring/FilteredColimits.lean | 33 +- Mathlib/Algebra/Category/Ring/Limits.lean | 88 +-- .../Homology/DerivedCategory/Ext/Basic.lean | 2 +- .../Homology/HomologicalComplexBiprod.lean | 2 +- .../Homology/HomologicalComplexLimits.lean | 28 +- .../ShortComplex/ConcreteCategory.lean | 4 +- .../Homology/ShortComplex/ExactFunctor.lean | 46 +- .../Algebra/Homology/ShortComplex/Limits.lean | 14 +- .../ShortComplex/PreservesHomology.lean | 60 +- Mathlib/AlgebraicGeometry/AffineScheme.lean | 2 +- .../GammaSpecAdjunction.lean | 4 +- Mathlib/AlgebraicGeometry/Limits.lean | 24 +- Mathlib/AlgebraicGeometry/OpenImmersion.lean | 18 +- .../FundamentalGroupoid/Product.lean | 4 +- Mathlib/CategoryTheory/Abelian/Exact.lean | 14 +- .../Abelian/GrothendieckAxioms.lean | 18 +- Mathlib/CategoryTheory/Abelian/Injective.lean | 8 +- .../CategoryTheory/Abelian/Projective.lean | 8 +- Mathlib/CategoryTheory/Abelian/Transfer.lean | 8 +- Mathlib/CategoryTheory/Adhesive.lean | 8 +- Mathlib/CategoryTheory/Adjunction/Limits.lean | 97 ++- .../CategoryTheory/Category/Cat/Limit.lean | 2 +- .../CategoryTheory/ChosenFiniteProducts.lean | 24 +- .../ChosenFiniteProducts/FunctorCategory.lean | 4 +- Mathlib/CategoryTheory/Closed/Cartesian.lean | 2 +- Mathlib/CategoryTheory/Closed/Ideal.lean | 28 +- Mathlib/CategoryTheory/Closed/Monoidal.lean | 2 +- Mathlib/CategoryTheory/Extensive.lean | 20 +- Mathlib/CategoryTheory/Functor/Flat.lean | 78 +- Mathlib/CategoryTheory/Galois/Action.lean | 6 +- Mathlib/CategoryTheory/Galois/Basic.lean | 14 +- Mathlib/CategoryTheory/Galois/Examples.lean | 2 +- .../CategoryTheory/Galois/GaloisObjects.lean | 2 +- .../Galois/Prorepresentability.lean | 2 +- Mathlib/CategoryTheory/GlueData.lean | 2 +- .../CategoryTheory/GradedObject/Monoidal.lean | 1 + Mathlib/CategoryTheory/Limits/Comma.lean | 9 +- .../CategoryTheory/Limits/ConeCategory.lean | 10 +- Mathlib/CategoryTheory/Limits/Connected.lean | 9 +- .../Limits/Constructions/BinaryProducts.lean | 8 +- .../Limits/Constructions/Equalizers.lean | 9 +- .../FiniteProductsOfBinaryProducts.lean | 36 +- .../LimitsOfProductsAndEqualizers.lean | 47 +- Mathlib/CategoryTheory/Limits/Creates.lean | 79 +- .../CategoryTheory/Limits/ExactFunctor.lean | 36 +- .../FilteredColimitCommutesFiniteLimit.lean | 16 +- Mathlib/CategoryTheory/Limits/FintypeCat.lean | 12 +- .../Limits/FunctorCategory/Basic.lean | 83 ++- Mathlib/CategoryTheory/Limits/Over.lean | 8 +- .../Limits/Preserves/Basic.lean | 679 +++++++++++++----- .../Limits/Preserves/Filtered.lean | 78 +- .../Limits/Preserves/Finite.lean | 158 ++-- .../Limits/Preserves/FunctorCategory.lean | 63 +- .../Limits/Preserves/Limits.lean | 22 +- .../Limits/Preserves/Opposites.lean | 374 +++++----- .../Limits/Preserves/Presheaf.lean | 6 +- .../Preserves/Shapes/BinaryProducts.lean | 16 +- .../Limits/Preserves/Shapes/Biproducts.lean | 51 +- .../Limits/Preserves/Shapes/Equalizers.lean | 28 +- .../Limits/Preserves/Shapes/Kernels.lean | 28 +- .../Limits/Preserves/Shapes/Products.lean | 26 +- .../Limits/Preserves/Shapes/Pullbacks.lean | 46 +- .../Limits/Preserves/Shapes/Terminal.lean | 55 +- .../Limits/Preserves/Shapes/Zero.lean | 34 +- .../Limits/Preserves/Ulift.lean | 7 +- .../Limits/Preserves/Yoneda.lean | 2 +- Mathlib/CategoryTheory/Limits/Presheaf.lean | 19 +- .../Limits/Shapes/NormalMono/Basic.lean | 28 +- Mathlib/CategoryTheory/Limits/VanKampen.lean | 12 +- Mathlib/CategoryTheory/Limits/Yoneda.lean | 46 +- .../Localization/FiniteProducts.lean | 7 +- .../CategoryTheory/Monad/Comonadicity.lean | 10 +- Mathlib/CategoryTheory/Monad/Limits.lean | 54 +- Mathlib/CategoryTheory/Monad/Monadicity.lean | 10 +- .../Monoidal/Internal/Limits.lean | 4 +- .../CategoryTheory/Monoidal/Preadditive.lean | 8 +- .../Preadditive/AdditiveFunctor.lean | 10 +- .../Preadditive/Biproducts.lean | 113 +-- .../CategoryTheory/Preadditive/Injective.lean | 4 +- .../CategoryTheory/Preadditive/LeftExact.lean | 73 +- .../Preadditive/Projective.lean | 4 +- .../Preadditive/Yoneda/Limits.lean | 13 +- Mathlib/CategoryTheory/Sites/Abelian.lean | 2 +- .../Sites/Coherent/ExtensiveSheaves.lean | 19 +- .../Sites/Coherent/LocallySurjective.lean | 2 +- .../Sites/Coherent/SheafComparison.lean | 26 +- Mathlib/CategoryTheory/Sites/Equivalence.lean | 2 +- Mathlib/CategoryTheory/Sites/LeftExact.lean | 80 ++- Mathlib/CategoryTheory/Sites/Preserves.lean | 23 +- Mathlib/CategoryTheory/Sites/Pullback.lean | 6 +- Mathlib/CategoryTheory/Sites/Sheaf.lean | 6 +- .../CategoryTheory/Sites/Sheafification.lean | 2 +- .../CategoryTheory/Triangulated/Functor.lean | 4 +- .../Triangulated/HomologicalFunctor.lean | 6 +- .../Condensed/Discrete/Characterization.lean | 8 +- Mathlib/Condensed/Discrete/Colimit.lean | 4 +- Mathlib/Condensed/Explicit.lean | 18 +- Mathlib/Condensed/Light/Explicit.lean | 6 +- Mathlib/Condensed/TopComparison.lean | 6 +- .../LocallyRingedSpace/HasColimits.lean | 8 +- .../Geometry/RingedSpace/OpenImmersion.lean | 58 +- .../PresheafedSpace/HasColimits.lean | 15 +- .../Geometry/RingedSpace/SheafedSpace.lean | 4 +- .../RepresentationTheory/Action/Basic.lean | 10 +- .../RepresentationTheory/Action/Limits.lean | 52 +- Mathlib/RepresentationTheory/Rep.lean | 4 +- .../Algebra/Category/ProfiniteGrp/Basic.lean | 2 +- .../Category/CompHausLike/Limits.lean | 10 +- .../Category/LightProfinite/Basic.lean | 4 +- .../Topology/Category/Profinite/Basic.lean | 4 +- .../Category/Stonean/Adjunctions.lean | 2 +- .../Category/TopCat/Limits/Basic.lean | 19 +- Mathlib/Topology/Category/TopCat/Yoneda.lean | 4 +- .../SheafCondition/PairwiseIntersections.lean | 4 +- 129 files changed, 2102 insertions(+), 1712 deletions(-) diff --git a/Mathlib/Algebra/Category/AlgebraCat/Limits.lean b/Mathlib/Algebra/Category/AlgebraCat/Limits.lean index 44fe90192172a..94f312fee5673 100644 --- a/Mathlib/Algebra/Category/AlgebraCat/Limits.lean +++ b/Mathlib/Algebra/Category/AlgebraCat/Limits.lean @@ -155,41 +155,41 @@ instance hasLimits : HasLimits (AlgebraCat.{w} R) := /-- The forgetful functor from R-algebras to rings preserves all limits. -/ -instance forget₂RingPreservesLimitsOfSize [UnivLE.{v, w}] : +instance forget₂Ring_preservesLimitsOfSize [UnivLE.{v, w}] : PreservesLimitsOfSize.{t, v} (forget₂ (AlgebraCat.{w} R) RingCat.{w}) where preservesLimitsOfShape := { preservesLimit := fun {K} ↦ - preservesLimitOfPreservesLimitCone (limitConeIsLimit K) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit K) (RingCat.limitConeIsLimit.{v, w} (_ ⋙ forget₂ (AlgebraCat.{w} R) RingCat.{w})) } -instance forget₂RingPreservesLimits : PreservesLimits (forget₂ (AlgebraCat R) RingCat.{w}) := - AlgebraCat.forget₂RingPreservesLimitsOfSize.{w, w} +instance forget₂Ring_preservesLimits : PreservesLimits (forget₂ (AlgebraCat R) RingCat.{w}) := + AlgebraCat.forget₂Ring_preservesLimitsOfSize.{w, w} /-- The forgetful functor from R-algebras to R-modules preserves all limits. -/ -instance forget₂ModulePreservesLimitsOfSize [UnivLE.{v, w}] : PreservesLimitsOfSize.{t, v} +instance forget₂Module_preservesLimitsOfSize [UnivLE.{v, w}] : PreservesLimitsOfSize.{t, v} (forget₂ (AlgebraCat.{w} R) (ModuleCat.{w} R)) where preservesLimitsOfShape := { preservesLimit := fun {K} ↦ - preservesLimitOfPreservesLimitCone (limitConeIsLimit K) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit K) (ModuleCat.HasLimits.limitConeIsLimit (K ⋙ forget₂ (AlgebraCat.{w} R) (ModuleCat.{w} R))) } -instance forget₂ModulePreservesLimits : +instance forget₂Module_preservesLimits : PreservesLimits (forget₂ (AlgebraCat R) (ModuleCat.{w} R)) := - AlgebraCat.forget₂ModulePreservesLimitsOfSize.{w, w} + AlgebraCat.forget₂Module_preservesLimitsOfSize.{w, w} /-- The forgetful functor from R-algebras to types preserves all limits. -/ -instance forgetPreservesLimitsOfSize [UnivLE.{v, w}] : +instance forget_preservesLimitsOfSize [UnivLE.{v, w}] : PreservesLimitsOfSize.{t, v} (forget (AlgebraCat.{w} R)) where preservesLimitsOfShape := { preservesLimit := fun {K} ↦ - preservesLimitOfPreservesLimitCone (limitConeIsLimit K) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit K) (Types.Small.limitConeIsLimit.{v} (K ⋙ forget _)) } -instance forgetPreservesLimits : PreservesLimits (forget (AlgebraCat.{w} R)) := - AlgebraCat.forgetPreservesLimitsOfSize.{w, w} +instance forget_preservesLimits : PreservesLimits (forget (AlgebraCat.{w} R)) := + AlgebraCat.forget_preservesLimitsOfSize.{w, w} end AlgebraCat diff --git a/Mathlib/Algebra/Category/Grp/AB5.lean b/Mathlib/Algebra/Category/Grp/AB5.lean index 2edc3d04999e0..a5230f8bf90ff 100644 --- a/Mathlib/Algebra/Category/Grp/AB5.lean +++ b/Mathlib/Algebra/Category/Grp/AB5.lean @@ -27,7 +27,7 @@ variable {J : Type u} [SmallCategory J] [IsFiltered J] noncomputable instance : (colim (J := J) (C := AddCommGrp.{u})).PreservesHomology := - Functor.preservesHomologyOfMapExact _ (fun S hS ↦ by + Functor.preservesHomology_of_map_exact _ (fun S hS ↦ by replace hS := fun j => hS.map ((evaluation _ _).obj j) simp only [ShortComplex.ab_exact_iff_ker_le_range] at hS ⊢ intro x (hx : _ = _) @@ -45,7 +45,7 @@ noncomputable instance : noncomputable instance : PreservesFiniteLimits <| colim (J := J) (C := AddCommGrp.{u}) := by - apply Functor.preservesFiniteLimitsOfPreservesHomology + apply Functor.preservesFiniteLimits_of_preservesHomology instance : HasFilteredColimits (AddCommGrp.{u}) where HasColimitsOfShape := inferInstance diff --git a/Mathlib/Algebra/Category/Grp/FilteredColimits.lean b/Mathlib/Algebra/Category/Grp/FilteredColimits.lean index b79cc51e0f2ed..a31573223addd 100644 --- a/Mathlib/Algebra/Category/Grp/FilteredColimits.lean +++ b/Mathlib/Algebra/Category/Grp/FilteredColimits.lean @@ -128,18 +128,18 @@ def colimitCoconeIsColimit : IsColimit (colimitCocone.{v, u} F) where ((forget Grp).mapCocone t) _ fun j => funext fun x => DFunLike.congr_fun (h j) x -@[to_additive forget₂AddMonPreservesFilteredColimits] -noncomputable instance forget₂MonPreservesFilteredColimits : +@[to_additive forget₂AddMon_preservesFilteredColimits] +noncomputable instance forget₂Mon_preservesFilteredColimits : PreservesFilteredColimits.{u} (forget₂ Grp.{u} MonCat.{u}) where preserves_filtered_colimits x hx1 _ := letI : Category.{u, u} x := hx1 - ⟨fun {F} => preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit.{u, u} F) + ⟨fun {F} => preservesColimit_of_preserves_colimit_cocone (colimitCoconeIsColimit.{u, u} F) (MonCat.FilteredColimits.colimitCoconeIsColimit.{u, u} _)⟩ @[to_additive] -noncomputable instance forgetPreservesFilteredColimits : +noncomputable instance forget_preservesFilteredColimits : PreservesFilteredColimits (forget Grp.{u}) := - Limits.compPreservesFilteredColimits (forget₂ Grp MonCat) (forget MonCat.{u}) + Limits.comp_preservesFilteredColimits (forget₂ Grp MonCat) (forget MonCat.{u}) end @@ -198,19 +198,19 @@ def colimitCoconeIsColimit : IsColimit (colimitCocone.{v, u} F) where ((forget CommGrp).mapCocone t) _ fun j => funext fun x => DFunLike.congr_fun (h j) x @[to_additive] -noncomputable instance forget₂GroupPreservesFilteredColimits : +noncomputable instance forget₂Group_preservesFilteredColimits : PreservesFilteredColimits (forget₂ CommGrp Grp.{u}) where preserves_filtered_colimits J hJ1 _ := letI : Category J := hJ1 { preservesColimit := fun {F} => - preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit.{u, u} F) + preservesColimit_of_preserves_colimit_cocone (colimitCoconeIsColimit.{u, u} F) (Grp.FilteredColimits.colimitCoconeIsColimit.{u, u} (F ⋙ forget₂ CommGrp Grp.{u})) } @[to_additive] -noncomputable instance forgetPreservesFilteredColimits : +noncomputable instance forget_preservesFilteredColimits : PreservesFilteredColimits (forget CommGrp.{u}) := - Limits.compPreservesFilteredColimits (forget₂ CommGrp Grp) (forget Grp.{u}) + Limits.comp_preservesFilteredColimits (forget₂ CommGrp Grp) (forget Grp.{u}) end diff --git a/Mathlib/Algebra/Category/Grp/Limits.lean b/Mathlib/Algebra/Category/Grp/Limits.lean index 87ad275a4593f..f9226f51a1742 100644 --- a/Mathlib/Algebra/Category/Grp/Limits.lean +++ b/Mathlib/Algebra/Category/Grp/Limits.lean @@ -161,21 +161,21 @@ This means the underlying monoid of a limit can be computed as a limit in the ca This means the underlying additive monoid of a limit can be computed as a limit in the category of additive monoids.", to_additive_relevant_arg 2] -noncomputable instance forget₂MonPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget₂Mon_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget₂ Grp.{u} MonCat.{u}) where preservesLimitsOfShape {J _} := { } @[to_additive] -noncomputable instance forget₂MonPreservesLimits : +instance forget₂Mon_preservesLimits : PreservesLimits (forget₂ Grp.{u} MonCat.{u}) := - Grp.forget₂MonPreservesLimitsOfSize.{u, u} + Grp.forget₂Mon_preservesLimitsOfSize.{u, u} /-- If `J` is `u`-small, the forgetful functor from `Grp.{u}` preserves limits of shape `J`. -/ @[to_additive "If `J` is `u`-small, the forgetful functor from `AddGrp.{u}`\n preserves limits of shape `J`."] -noncomputable instance forgetPreservesLimitsOfShape [Small.{u} J] : +instance forget_preservesLimitsOfShape [Small.{u} J] : PreservesLimitsOfShape J (forget Grp.{u}) where - preservesLimit {F} := preservesLimitOfPreservesLimitCone (limitConeIsLimit F) + preservesLimit {F} := preservesLimit_of_preserves_limit_cone (limitConeIsLimit F) (Types.Small.limitConeIsLimit (F ⋙ forget _)) /-- The forgetful functor from groups to types preserves all limits. @@ -186,12 +186,12 @@ This means the underlying type of a limit can be computed as a limit in the cate This means the underlying type of a limit can be computed as a limit in the category of types.", to_additive_relevant_arg 2] -noncomputable instance forgetPreservesLimitsOfSize : +instance forget_preservesLimitsOfSize : PreservesLimitsOfSize.{w, v} (forget Grp.{u}) := inferInstance @[to_additive] -noncomputable instance forgetPreservesLimits : PreservesLimits (forget Grp.{u}) := - Grp.forgetPreservesLimitsOfSize.{u, u} +instance forget_preservesLimits : PreservesLimits (forget Grp.{u}) := + Grp.forget_preservesLimitsOfSize.{u, u} end Grp @@ -307,18 +307,18 @@ instance hasLimits : HasLimits CommGrp.{u} := CommGrp.hasLimitsOfSize.{u, u} @[to_additive] -noncomputable instance forget₂GroupPreservesLimit : +instance forget₂Group_preservesLimit : PreservesLimit F (forget₂ CommGrp.{u} Grp.{u}) where - preserves {c} hc := by + preserves {c} hc := ⟨by have : HasLimit (F ⋙ forget₂ CommGrp Grp) := by rw [Grp.hasLimit_iff_small_sections] change Small.{u} (F ⋙ forget CommGrp).sections rw [← CommGrp.hasLimit_iff_small_sections] exact ⟨_, hc⟩ - exact isLimitOfPreserves _ hc + exact isLimitOfPreserves _ hc⟩ @[to_additive] -noncomputable instance forget₂GroupPreservesLimitsOfShape : +instance forget₂Group_preservesLimitsOfShape : PreservesLimitsOfShape J (forget₂ CommGrp.{u} Grp.{u}) where /-- The forgetful functor from commutative groups to groups preserves all limits. @@ -330,19 +330,19 @@ of groups.) (That is, the underlying group could have been computed instead as limits in the category of additive groups.)", to_additive_relevant_arg 2] -noncomputable instance forget₂GroupPreservesLimitsOfSize : +instance forget₂Group_preservesLimitsOfSize : PreservesLimitsOfSize.{w, v} (forget₂ CommGrp.{u} Grp.{u}) where @[to_additive] -noncomputable instance forget₂GroupPreservesLimits : +instance forget₂Group_preservesLimits : PreservesLimits (forget₂ CommGrp Grp.{u}) := - CommGrp.forget₂GroupPreservesLimitsOfSize.{u, u} + CommGrp.forget₂Group_preservesLimitsOfSize.{u, u} /-- An auxiliary declaration to speed up typechecking. -/ -@[to_additive AddCommGrp.forget₂AddCommMonPreservesLimitsAux +@[to_additive AddCommGrp.forget₂AddCommMon_preservesLimitsAux "An auxiliary declaration to speed up typechecking."] -noncomputable def forget₂CommMonPreservesLimitsAux +noncomputable def forget₂CommMon_preservesLimitsAux [Small.{u} (F ⋙ forget CommGrp).sections] : IsLimit ((forget₂ CommGrp.{u} CommMonCat.{u}).mapCone (limitCone.{v, u} F)) := letI : Small.{u} (Functor.sections ((F ⋙ forget₂ _ CommMonCat) ⋙ forget CommMonCat)) := @@ -351,23 +351,23 @@ noncomputable def forget₂CommMonPreservesLimitsAux /-- If `J` is `u`-small, the forgetful functor from `CommGrp.{u}` to `CommMonCat.{u}` preserves limits of shape `J`. -/ -@[to_additive AddCommGrp.forget₂AddCommMonPreservesLimitsOfShape +@[to_additive AddCommGrp.forget₂AddCommMon_preservesLimitsOfShape "If `J` is `u`-small, the forgetful functor from `AddCommGrp.{u}` to `AddCommMonCat.{u}` preserves limits of shape `J`."] -noncomputable instance forget₂CommMonPreservesLimitsOfShape [Small.{u} J] : +instance forget₂CommMon_preservesLimitsOfShape [Small.{u} J] : PreservesLimitsOfShape J (forget₂ CommGrp.{u} CommMonCat.{u}) where - preservesLimit {F} := preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F) - (forget₂CommMonPreservesLimitsAux.{v, u} F) + preservesLimit {F} := preservesLimit_of_preserves_limit_cone (limitConeIsLimit.{v, u} F) + (forget₂CommMon_preservesLimitsAux.{v, u} F) /-- The forgetful functor from commutative groups to commutative monoids preserves all limits. (That is, the underlying commutative monoids could have been computed instead as limits in the category of commutative monoids.) -/ -@[to_additive AddCommGrp.forget₂AddCommMonPreservesLimitsOfSize +@[to_additive AddCommGrp.forget₂AddCommMon_preservesLimitsOfSize "The forgetful functor from additive commutative groups to additive commutative monoids preserves all limits. (That is, the underlying additive commutative monoids could have been computed instead as limits in the category of additive commutative monoids.)"] -noncomputable instance forget₂CommMonPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget₂CommMon_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget₂ CommGrp CommMonCat.{u}) where preservesLimitsOfShape := { } @@ -375,9 +375,9 @@ noncomputable instance forget₂CommMonPreservesLimitsOfSize [UnivLE.{v, u}] : shape `J`. -/ @[to_additive "If `J` is `u`-small, the forgetful functor from `AddCommGrp.{u}`\n preserves limits of shape `J`."] -noncomputable instance forgetPreservesLimitsOfShape [Small.{u} J] : +instance forget_preservesLimitsOfShape [Small.{u} J] : PreservesLimitsOfShape J (forget CommGrp.{u}) where - preservesLimit {F} := preservesLimitOfPreservesLimitCone (limitConeIsLimit F) + preservesLimit {F} := preservesLimit_of_preserves_limit_cone (limitConeIsLimit F) (Types.Small.limitConeIsLimit (F ⋙ forget _)) /-- The forgetful functor from commutative groups to types preserves all limits. (That is, the @@ -387,16 +387,16 @@ underlying types could have been computed instead as limits in the category of t "The forgetful functor from additive commutative groups to types preserves all limits. (That is, the underlying types could have been computed instead as limits in the category of types.)"] -noncomputable instance forgetPreservesLimitsOfSize : +instance forget_preservesLimitsOfSize : PreservesLimitsOfSize.{w, v} (forget CommGrp.{u}) := inferInstance -noncomputable instance _root_.AddCommGrp.forgetPreservesLimits : +noncomputable instance _root_.AddCommGrp.forget_preservesLimits : PreservesLimits (forget AddCommGrp.{u}) := - AddCommGrp.forgetPreservesLimitsOfSize.{u, u} + AddCommGrp.forget_preservesLimitsOfSize.{u, u} @[to_additive existing] -noncomputable instance forgetPreservesLimits : PreservesLimits (forget CommGrp.{u}) := - CommGrp.forgetPreservesLimitsOfSize.{u, u} +noncomputable instance forget_preservesLimits : PreservesLimits (forget CommGrp.{u}) := + CommGrp.forget_preservesLimitsOfSize.{u, u} -- Verify we can form limits indexed over smaller categories. example (f : ℕ → AddCommGrp) : HasProduct f := by infer_instance diff --git a/Mathlib/Algebra/Category/ModuleCat/Abelian.lean b/Mathlib/Algebra/Category/ModuleCat/Abelian.lean index 401fb715425c1..fcb99e727a6c4 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Abelian.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Abelian.lean @@ -82,19 +82,19 @@ instance : HasLimitsOfSize.{v,v} (ModuleCatMax.{v, w} R) := /- We need to put this in this weird spot because we need to know that the category of modules is balanced. -/ -instance forgetReflectsLimitsOfSize : +instance forget_reflectsLimitsOfSize : ReflectsLimitsOfSize.{v, v} (forget (ModuleCatMax.{v, w} R)) := - reflectsLimitsOfReflectsIsomorphisms + reflectsLimits_of_reflectsIsomorphisms -instance forget₂ReflectsLimitsOfSize : +instance forget₂_reflectsLimitsOfSize : ReflectsLimitsOfSize.{v, v} (forget₂ (ModuleCatMax.{v, w} R) AddCommGrp.{max v w}) := - reflectsLimitsOfReflectsIsomorphisms + reflectsLimits_of_reflectsIsomorphisms -instance forgetReflectsLimits : ReflectsLimits (forget (ModuleCat.{v} R)) := - ModuleCat.forgetReflectsLimitsOfSize.{v, v} +instance forget_reflectsLimits : ReflectsLimits (forget (ModuleCat.{v} R)) := + ModuleCat.forget_reflectsLimitsOfSize.{v, v} -instance forget₂ReflectsLimits : ReflectsLimits (forget₂ (ModuleCat.{v} R) AddCommGrp.{v}) := - ModuleCat.forget₂ReflectsLimitsOfSize.{v, v} +instance forget₂_reflectsLimits : ReflectsLimits (forget₂ (ModuleCat.{v} R) AddCommGrp.{v}) := + ModuleCat.forget₂_reflectsLimitsOfSize.{v, v} end ReflectsLimits diff --git a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean index 254e152bf938a..4bbcee39a959e 100644 --- a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean +++ b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean @@ -827,26 +827,25 @@ instance {R : Type u₁} {S : Type u₂} [CommRing R] [CommRing S] (f : R →+* (restrictScalars.{max u₂ w, u₁, u₂} f).IsRightAdjoint := (extendRestrictScalarsAdj f).isRightAdjoint -noncomputable instance preservesLimitRestrictScalars +noncomputable instance preservesLimit_restrictScalars {R : Type*} {S : Type*} [Ring R] [Ring S] (f : R →+* S) {J : Type*} [Category J] (F : J ⥤ ModuleCat.{v} S) [Small.{v} (F ⋙ forget _).sections] : PreservesLimit F (restrictScalars f) := - ⟨fun {c} hc => by + ⟨fun {c} hc => ⟨by have hc' := isLimitOfPreserves (forget₂ _ AddCommGrp) hc - exact isLimitOfReflects (forget₂ _ AddCommGrp) hc'⟩ + exact isLimitOfReflects (forget₂ _ AddCommGrp) hc'⟩⟩ -instance preservesColimitRestrictScalars {R S : Type*} [Ring R] [Ring S] +instance preservesColimit_restrictScalars {R S : Type*} [Ring R] [Ring S] (f : R →+* S) {J : Type*} [Category J] (F : J ⥤ ModuleCat.{v} S) [HasColimit (F ⋙ forget₂ _ AddCommGrp)] : PreservesColimit F (ModuleCat.restrictScalars.{v} f) := by have : HasColimit ((F ⋙ restrictScalars f) ⋙ forget₂ (ModuleCat R) AddCommGrp) := inferInstanceAs (HasColimit (F ⋙ forget₂ _ AddCommGrp)) - apply preservesColimitOfPreservesColimitCocone (HasColimit.isColimitColimitCocone F) + apply preservesColimit_of_preserves_colimit_cocone (HasColimit.isColimitColimitCocone F) apply isColimitOfReflects (forget₂ _ AddCommGrp) apply isColimitOfPreserves (forget₂ (ModuleCat.{v} S) AddCommGrp.{v}) exact HasColimit.isColimitColimitCocone F - end ModuleCat end ModuleCat diff --git a/Mathlib/Algebra/Category/ModuleCat/Colimits.lean b/Mathlib/Algebra/Category/ModuleCat/Colimits.lean index 726d06d08e305..358c2eb1ea0fc 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Colimits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Colimits.lean @@ -97,11 +97,11 @@ noncomputable def isColimitColimitCocone : IsColimit (colimitCocone F) where instance : HasColimit F := ⟨_, isColimitColimitCocone F⟩ noncomputable instance : PreservesColimit F (forget₂ _ AddCommGrp) := - preservesColimitOfPreservesColimitCocone (isColimitColimitCocone F) (colimit.isColimit _) + preservesColimit_of_preserves_colimit_cocone (isColimitColimitCocone F) (colimit.isColimit _) noncomputable instance reflectsColimit : ReflectsColimit F (forget₂ (ModuleCat.{w'} R) AddCommGrp) := - reflectsColimitOfReflectsIsomorphisms _ _ + reflectsColimit_of_reflectsIsomorphisms _ _ end HasColimit diff --git a/Mathlib/Algebra/Category/ModuleCat/FilteredColimits.lean b/Mathlib/Algebra/Category/ModuleCat/FilteredColimits.lean index 4ab0442cf4f77..553ce84f4002e 100644 --- a/Mathlib/Algebra/Category/ModuleCat/FilteredColimits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/FilteredColimits.lean @@ -170,22 +170,22 @@ def colimitCoconeIsColimit : IsColimit (colimitCocone F) where (Types.TypeMax.colimitCoconeIsColimit (F ⋙ forget (ModuleCat R))).uniq ((forget (ModuleCat R)).mapCocone t) _ fun j => funext fun x => LinearMap.congr_fun (h j) x -instance forget₂AddCommGroupPreservesFilteredColimits : +instance forget₂AddCommGroup_preservesFilteredColimits : PreservesFilteredColimits (forget₂ (ModuleCat.{u} R) AddCommGrp.{u}) where preserves_filtered_colimits J _ _ := { -- Porting note: without the curly braces for `F` -- here we get a confusing error message about universes. preservesColimit := fun {F : J ⥤ ModuleCat.{u} R} => - preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit F) + preservesColimit_of_preserves_colimit_cocone (colimitCoconeIsColimit F) (AddCommGrp.FilteredColimits.colimitCoconeIsColimit (F ⋙ forget₂ (ModuleCat.{u} R) AddCommGrp.{u})) } -instance forgetPreservesFilteredColimits : PreservesFilteredColimits (forget (ModuleCat.{u} R)) := - Limits.compPreservesFilteredColimits (forget₂ (ModuleCat R) AddCommGrp) +instance forget_preservesFilteredColimits : PreservesFilteredColimits (forget (ModuleCat.{u} R)) := + Limits.comp_preservesFilteredColimits (forget₂ (ModuleCat R) AddCommGrp) (forget AddCommGrp) -instance forgetReflectsFilteredColimits : ReflectsFilteredColimits (forget (ModuleCat.{u} R)) where - reflects_filtered_colimits _ := { reflectsColimit := reflectsColimitOfReflectsIsomorphisms _ _ } +instance forget_reflectsFilteredColimits : ReflectsFilteredColimits (forget (ModuleCat.{u} R)) where + reflects_filtered_colimits _ := { reflectsColimit := reflectsColimit_of_reflectsIsomorphisms _ _ } end diff --git a/Mathlib/Algebra/Category/ModuleCat/Limits.lean b/Mathlib/Algebra/Category/ModuleCat/Limits.lean index dab88fdebc6cf..862547dd00ba8 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Limits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Limits.lean @@ -149,7 +149,7 @@ instance (priority := high) hasLimits' : HasLimits (ModuleCat.{u} R) := /-- An auxiliary declaration to speed up typechecking. -/ -def forget₂AddCommGroupPreservesLimitsAux : +def forget₂AddCommGroup_preservesLimitsAux : IsLimit ((forget₂ (ModuleCat R) AddCommGrp).mapCone (limitCone F)) := letI : Small.{w} (Functor.sections ((F ⋙ forget₂ _ AddCommGrp) ⋙ forget _)) := inferInstanceAs <| Small.{w} (Functor.sections (F ⋙ forget (ModuleCat R))) @@ -157,48 +157,48 @@ def forget₂AddCommGroupPreservesLimitsAux : (F ⋙ forget₂ (ModuleCat.{w} R) _ : J ⥤ AddCommGrp.{w}) /-- The forgetful functor from R-modules to abelian groups preserves all limits. -/ -instance forget₂AddCommGroupPreservesLimit : +instance forget₂AddCommGroup_preservesLimit : PreservesLimit F (forget₂ (ModuleCat R) AddCommGrp) := - preservesLimitOfPreservesLimitCone (limitConeIsLimit F) - (forget₂AddCommGroupPreservesLimitsAux F) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit F) + (forget₂AddCommGroup_preservesLimitsAux F) /-- The forgetful functor from R-modules to abelian groups preserves all limits. -/ -instance forget₂AddCommGroupPreservesLimitsOfSize [UnivLE.{v, w}] : +instance forget₂AddCommGroup_preservesLimitsOfSize [UnivLE.{v, w}] : PreservesLimitsOfSize.{t, v} (forget₂ (ModuleCat.{w} R) AddCommGrp.{w}) where preservesLimitsOfShape := { preservesLimit := inferInstance } -instance forget₂AddCommGroupPreservesLimits : +instance forget₂AddCommGroup_preservesLimits : PreservesLimits (forget₂ (ModuleCat R) AddCommGrp.{w}) := - ModuleCat.forget₂AddCommGroupPreservesLimitsOfSize.{w, w} + ModuleCat.forget₂AddCommGroup_preservesLimitsOfSize.{w, w} /-- The forgetful functor from R-modules to types preserves all limits. -/ -instance forgetPreservesLimitsOfSize [UnivLE.{v, w}] : +instance forget_preservesLimitsOfSize [UnivLE.{v, w}] : PreservesLimitsOfSize.{t, v} (forget (ModuleCat.{w} R)) where preservesLimitsOfShape := - { preservesLimit := fun {K} ↦ preservesLimitOfPreservesLimitCone (limitConeIsLimit K) + { preservesLimit := fun {K} ↦ preservesLimit_of_preserves_limit_cone (limitConeIsLimit K) (Types.Small.limitConeIsLimit.{v} (_ ⋙ forget _)) } -instance forgetPreservesLimits : PreservesLimits (forget (ModuleCat.{w} R)) := - ModuleCat.forgetPreservesLimitsOfSize.{w, w} +instance forget_preservesLimits : PreservesLimits (forget (ModuleCat.{w} R)) := + ModuleCat.forget_preservesLimitsOfSize.{w, w} end -instance forget₂AddCommGroupReflectsLimit : +instance forget₂AddCommGroup_reflectsLimit : ReflectsLimit F (forget₂ (ModuleCat.{w} R) AddCommGrp) where - reflects {c} hc := by + reflects {c} hc := ⟨by have : HasLimit (F ⋙ forget₂ (ModuleCat R) AddCommGrp) := ⟨_, hc⟩ have : Small.{w} (Functor.sections (F ⋙ forget (ModuleCat R))) := by simpa only [AddCommGrp.hasLimit_iff_small_sections] using this - have := reflectsLimitOfReflectsIsomorphisms F (forget₂ (ModuleCat R) AddCommGrp) - exact isLimitOfReflects _ hc + have := reflectsLimit_of_reflectsIsomorphisms F (forget₂ (ModuleCat R) AddCommGrp) + exact isLimitOfReflects _ hc⟩ -instance forget₂AddCommGroupReflectsLimitOfShape : +instance forget₂AddCommGroup_reflectsLimitOfShape : ReflectsLimitsOfShape J (forget₂ (ModuleCat.{w} R) AddCommGrp) where -instance forget₂AddCommGroupReflectsLimitOfSize : +instance forget₂AddCommGroup_reflectsLimitOfSize : ReflectsLimitsOfSize.{t, v} (forget₂ (ModuleCat.{w} R) AddCommGrp) where section DirectLimit diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Colimits.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Colimits.lean index f58ed17c074ed..fa22a86db7365 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Colimits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Colimits.lean @@ -106,16 +106,16 @@ noncomputable def isColimitColimitCocone : IsColimit (colimitCocone F) := instance hasColimit : HasColimit F := ⟨_, isColimitColimitCocone F⟩ -noncomputable instance evaluationPreservesColimit (X : Cᵒᵖ) : +instance evaluation_preservesColimit (X : Cᵒᵖ) : PreservesColimit F (evaluation R X) := - preservesColimitOfPreservesColimitCocone (isColimitColimitCocone F) (colimit.isColimit _) + preservesColimit_of_preserves_colimit_cocone (isColimitColimitCocone F) (colimit.isColimit _) variable [∀ X, PreservesColimit F (evaluation R X ⋙ forget₂ (ModuleCat (R.obj X)) AddCommGrp)] -noncomputable instance toPresheafPreservesColimit : +instance toPresheaf_preservesColimit : PreservesColimit F (toPresheaf R) := - preservesColimitOfPreservesColimitCocone (isColimitColimitCocone F) + preservesColimit_of_preserves_colimit_cocone (isColimitColimitCocone F) (Limits.evaluationJointlyReflectsColimits _ (fun X => isColimitOfPreserves (evaluation R X ⋙ forget₂ _ AddCommGrp) (isColimitColimitCocone F))) @@ -130,10 +130,10 @@ variable [HasColimitsOfShape J AddCommGrp.{v}] instance hasColimitsOfShape : HasColimitsOfShape J (PresheafOfModules.{v} R) where -noncomputable instance evaluationPreservesColimitsOfShape (X : Cᵒᵖ) : +noncomputable instance evaluation_preservesColimitsOfShape (X : Cᵒᵖ) : PreservesColimitsOfShape J (evaluation R X : PresheafOfModules.{v} R ⥤ _) where -noncomputable instance toPresheafPreservesColimitsOfShape : +noncomputable instance toPresheaf_preservesColimitsOfShape : PreservesColimitsOfShape J (toPresheaf.{v} R) where end HasColimitsOfShape @@ -143,10 +143,10 @@ namespace Finite instance hasFiniteColimits : HasFiniteColimits (PresheafOfModules.{v} R) := ⟨fun _ => inferInstance⟩ -noncomputable instance evaluationPreservesFiniteColimits (X : Cᵒᵖ) : +noncomputable instance evaluation_preservesFiniteColimits (X : Cᵒᵖ) : PreservesFiniteColimits (evaluation.{v} R X) where -noncomputable instance toPresheafPreservesFiniteColimits : +noncomputable instance toPresheaf_preservesFiniteColimits : PreservesFiniteColimits (toPresheaf R) where end Finite diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Limits.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Limits.lean index 127f29b50a89e..12b4f9fc866eb 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Limits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Limits.lean @@ -118,13 +118,13 @@ noncomputable def isLimitLimitCone : IsLimit (limitCone F) := instance hasLimit : HasLimit F := ⟨_, isLimitLimitCone F⟩ -noncomputable instance evaluationPreservesLimit (X : Cᵒᵖ) : +noncomputable instance evaluation_preservesLimit (X : Cᵒᵖ) : PreservesLimit F (evaluation R X) := - preservesLimitOfPreservesLimitCone (isLimitLimitCone F) (limit.isLimit _) + preservesLimit_of_preserves_limit_cone (isLimitLimitCone F) (limit.isLimit _) -noncomputable instance toPresheafPreservesLimit : +noncomputable instance toPresheaf_preservesLimit : PreservesLimit F (toPresheaf R) := - preservesLimitOfPreservesLimitCone (isLimitLimitCone F) + preservesLimit_of_preserves_limit_cone (isLimitLimitCone F) (Limits.evaluationJointlyReflectsLimits _ (fun X => isLimitOfPreserves (evaluation R X ⋙ forget₂ _ AddCommGrp) (isLimitLimitCone F))) @@ -139,10 +139,10 @@ variable [Small.{v} J] instance hasLimitsOfShape : HasLimitsOfShape J (PresheafOfModules.{v} R) where -noncomputable instance evaluationPreservesLimitsOfShape (X : Cᵒᵖ) : +noncomputable instance evaluation_preservesLimitsOfShape (X : Cᵒᵖ) : PreservesLimitsOfShape J (evaluation R X : PresheafOfModules.{v} R ⥤ _) where -noncomputable instance toPresheafPreservesLimitsOfShape : +noncomputable instance toPresheaf_preservesLimitsOfShape : PreservesLimitsOfShape J (toPresheaf.{v} R) where end Small @@ -152,10 +152,10 @@ section Finite instance hasFiniteLimits : HasFiniteLimits (PresheafOfModules.{v} R) := ⟨fun _ => inferInstance⟩ -noncomputable instance evaluationPreservesFiniteLimits (X : Cᵒᵖ) : +noncomputable instance evaluation_preservesFiniteLimits (X : Cᵒᵖ) : PreservesFiniteLimits (evaluation.{v} R X) where -noncomputable instance toPresheafPreservesFiniteLimits : +noncomputable instance toPresheaf_preservesFiniteLimits : PreservesFiniteLimits (toPresheaf.{v} R) where end Finite diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafification.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafification.lean index 55d42463d8a56..949298d7e1b84 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafification.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafification.lean @@ -144,7 +144,7 @@ variable [HasSheafify J AddCommGrp.{v}] noncomputable instance : PreservesFiniteLimits (sheafification.{v} α ⋙ SheafOfModules.toSheaf.{v} R) := - compPreservesFiniteLimits (toPresheaf.{v} R₀) (presheafToSheaf J AddCommGrp) + comp_preservesFiniteLimits (toPresheaf.{v} R₀) (presheafToSheaf J AddCommGrp) instance : (SheafOfModules.toSheaf.{v} R ⋙ sheafToPresheaf _ _).ReflectsIsomorphisms := inferInstanceAs (SheafOfModules.forget.{v} R ⋙ toPresheaf _).ReflectsIsomorphisms @@ -156,7 +156,7 @@ noncomputable instance : ReflectsFiniteLimits (SheafOfModules.toSheaf.{v} R) whe reflects _ _ _ := inferInstance noncomputable instance : PreservesFiniteLimits (sheafification.{v} α) := - preservesFiniteLimitsOfReflectsOfPreserves + preservesFiniteLimits_of_reflects_of_preserves (sheafification.{v} α) (SheafOfModules.toSheaf.{v} R) end diff --git a/Mathlib/Algebra/Category/ModuleCat/Sheaf/Limits.lean b/Mathlib/Algebra/Category/ModuleCat/Sheaf/Limits.lean index d82c0027f62ba..91404085a45b9 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Sheaf/Limits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Sheaf/Limits.lean @@ -111,9 +111,9 @@ noncomputable instance forgetPreservesLimitsOfSize : noncomputable instance : PreservesFiniteLimits (SheafOfModules.toSheaf.{v} R ⋙ sheafToPresheaf _ _) := - compPreservesFiniteLimits (SheafOfModules.forget.{v} R) (PresheafOfModules.toPresheaf R.val) + comp_preservesFiniteLimits (SheafOfModules.forget.{v} R) (PresheafOfModules.toPresheaf R.val) noncomputable instance : PreservesFiniteLimits (SheafOfModules.toSheaf.{v} R) := - preservesFiniteLimitsOfReflectsOfPreserves _ (sheafToPresheaf _ _) + preservesFiniteLimits_of_reflects_of_preserves _ (sheafToPresheaf _ _) end SheafOfModules diff --git a/Mathlib/Algebra/Category/MonCat/FilteredColimits.lean b/Mathlib/Algebra/Category/MonCat/FilteredColimits.lean index db7946f579819..d12f9961dbc01 100644 --- a/Mathlib/Algebra/Category/MonCat/FilteredColimits.lean +++ b/Mathlib/Algebra/Category/MonCat/FilteredColimits.lean @@ -296,10 +296,10 @@ def colimitCoconeIsColimit : IsColimit (colimitCocone.{v, u} F) where fun j => funext fun x => DFunLike.congr_fun (i := MonCat.instFunLike _ _) (h j) x) y @[to_additive] -noncomputable instance forgetPreservesFilteredColimits : +instance forget_preservesFilteredColimits : PreservesFilteredColimits (forget MonCat.{u}) := ⟨fun J hJ1 _ => letI hJ1' : Category J := hJ1 - ⟨fun {F} => preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit.{u, u} F) + ⟨fun {F} => preservesColimit_of_preserves_colimit_cocone (colimitCoconeIsColimit.{u, u} F) (Types.TypeMax.colimitCoconeIsColimit (F ⋙ forget MonCat.{u}))⟩⟩ end @@ -369,16 +369,16 @@ def colimitCoconeIsColimit : IsColimit (colimitCocone.{v, u} F) where DFunLike.congr_fun (i := CommMonCat.instFunLike _ _) (h j) x @[to_additive forget₂AddMonPreservesFilteredColimits] -noncomputable instance forget₂MonPreservesFilteredColimits : +noncomputable instance forget₂Mon_preservesFilteredColimits : PreservesFilteredColimits (forget₂ CommMonCat MonCat.{u}) := ⟨fun J hJ1 _ => letI hJ3 : Category J := hJ1 - ⟨fun {F} => preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit.{u, u} F) + ⟨fun {F} => preservesColimit_of_preserves_colimit_cocone (colimitCoconeIsColimit.{u, u} F) (MonCat.FilteredColimits.colimitCoconeIsColimit (F ⋙ forget₂ CommMonCat MonCat.{u}))⟩⟩ @[to_additive] -noncomputable instance forgetPreservesFilteredColimits : +noncomputable instance forget_preservesFilteredColimits : PreservesFilteredColimits (forget CommMonCat.{u}) := - Limits.compPreservesFilteredColimits (forget₂ CommMonCat MonCat) (forget MonCat) + Limits.comp_preservesFilteredColimits (forget₂ CommMonCat MonCat) (forget MonCat) end diff --git a/Mathlib/Algebra/Category/MonCat/Limits.lean b/Mathlib/Algebra/Category/MonCat/Limits.lean index 933d00567f3f1..8145dc0c3be8e 100644 --- a/Mathlib/Algebra/Category/MonCat/Limits.lean +++ b/Mathlib/Algebra/Category/MonCat/Limits.lean @@ -137,9 +137,9 @@ instance hasLimits : HasLimits MonCat.{u} := /-- If `J` is `u`-small, the forgetful functor from `MonCat.{u}` preserves limits of shape `J`. -/ @[to_additive "If `J` is `u`-small, the forgetful functor from `AddMonCat.{u}`\n preserves limits of shape `J`."] -noncomputable instance forgetPreservesLimitsOfShape [Small.{u} J] : +noncomputable instance forget_preservesLimitsOfShape [Small.{u} J] : PreservesLimitsOfShape J (forget MonCat.{u}) where - preservesLimit {F} := preservesLimitOfPreservesLimitCone (limitConeIsLimit F) + preservesLimit {F} := preservesLimit_of_preserves_limit_cone (limitConeIsLimit F) (Types.Small.limitConeIsLimit (F ⋙ forget _)) /-- The forgetful functor from monoids to types preserves all limits. @@ -149,13 +149,13 @@ This means the underlying type of a limit can be computed as a limit in the cate "The forgetful functor from additive monoids to types preserves all limits.\n\n This means the underlying type of a limit can be computed as a limit in the category of types.", to_additive_relevant_arg 2] -noncomputable instance forgetPreservesLimitsOfSize [UnivLE.{v, u}] : +noncomputable instance forget_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget MonCat.{u}) where preservesLimitsOfShape := { } @[to_additive] -noncomputable instance forgetPreservesLimits : PreservesLimits (forget MonCat.{u}) := - MonCat.forgetPreservesLimitsOfSize.{u, u} +noncomputable instance forget_preservesLimits : PreservesLimits (forget MonCat.{u}) := + MonCat.forget_preservesLimitsOfSize.{u, u} end MonCat @@ -256,22 +256,22 @@ This means the underlying type of a limit can be computed as a limit in the cate This means the underlying type of a limit can be computed as a limit in the category of additive\n monoids.", to_additive_relevant_arg 2] -noncomputable instance forget₂MonPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget₂Mon_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget₂ CommMonCat.{u} MonCat.{u}) where preservesLimitsOfShape {J} 𝒥 := { } @[to_additive] -noncomputable instance forget₂MonPreservesLimits : +instance forget₂Mon_preservesLimits : PreservesLimits (forget₂ CommMonCat.{u} MonCat.{u}) := - CommMonCat.forget₂MonPreservesLimitsOfSize.{u, u} + CommMonCat.forget₂Mon_preservesLimitsOfSize.{u, u} /-- If `J` is `u`-small, the forgetful functor from `CommMonCat.{u}` preserves limits of shape `J`. -/ @[to_additive "If `J` is `u`-small, the forgetful functor from `AddCommMonCat.{u}`\n preserves limits of shape `J`."] -noncomputable instance forgetPreservesLimitsOfShape [Small.{u} J] : +instance forget_preservesLimitsOfShape [Small.{u} J] : PreservesLimitsOfShape J (forget CommMonCat.{u}) where - preservesLimit {F} := preservesLimitOfPreservesLimitCone (limitConeIsLimit F) + preservesLimit {F} := preservesLimit_of_preserves_limit_cone (limitConeIsLimit F) (Types.Small.limitConeIsLimit (F ⋙ forget _)) /-- The forgetful functor from commutative monoids to types preserves all limits. @@ -280,16 +280,16 @@ This means the underlying type of a limit can be computed as a limit in the cate @[to_additive "The forgetful functor from additive commutative monoids to types preserves all\n limits.\n\n This means the underlying type of a limit can be computed as a limit in the category of types."] -noncomputable instance forgetPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{v, v} (forget CommMonCat.{u}) where preservesLimitsOfShape {_} _ := { } -noncomputable instance _root_.AddCommMonCat.forgetPreservesLimits : +instance _root_.AddCommMonCat.forget_preservesLimits : PreservesLimits (forget AddCommMonCat.{u}) := - AddCommMonCat.forgetPreservesLimitsOfSize.{u, u} + AddCommMonCat.forget_preservesLimitsOfSize.{u, u} @[to_additive existing] -noncomputable instance forgetPreservesLimits : PreservesLimits (forget CommMonCat.{u}) := - CommMonCat.forgetPreservesLimitsOfSize.{u, u} +instance forget_preservesLimits : PreservesLimits (forget CommMonCat.{u}) := + CommMonCat.forget_preservesLimitsOfSize.{u, u} end CommMonCat diff --git a/Mathlib/Algebra/Category/Ring/FilteredColimits.lean b/Mathlib/Algebra/Category/Ring/FilteredColimits.lean index 9b383c982cbde..197ff8afa8f5c 100644 --- a/Mathlib/Algebra/Category/Ring/FilteredColimits.lean +++ b/Mathlib/Algebra/Category/Ring/FilteredColimits.lean @@ -137,16 +137,16 @@ def colimitCoconeIsColimit : IsColimit <| colimitCocone.{v, u} F where (Types.TypeMax.colimitCoconeIsColimit (F ⋙ forget SemiRingCat)).uniq ((forget SemiRingCat).mapCocone t) _ fun j => funext fun x => RingHom.congr_fun (h j) x -instance forget₂MonPreservesFilteredColimits : +instance forget₂Mon_preservesFilteredColimits : PreservesFilteredColimits (forget₂ SemiRingCat MonCat.{u}) where preserves_filtered_colimits {J hJ1 _} := letI : Category J := hJ1 { preservesColimit := fun {F} => - preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit.{u, u} F) + preservesColimit_of_preserves_colimit_cocone (colimitCoconeIsColimit.{u, u} F) (MonCat.FilteredColimits.colimitCoconeIsColimit (F ⋙ forget₂ SemiRingCat MonCat.{u})) } -instance forgetPreservesFilteredColimits : PreservesFilteredColimits (forget SemiRingCat.{u}) := - Limits.compPreservesFilteredColimits (forget₂ SemiRingCat MonCat) (forget MonCat.{u}) +instance forget_preservesFilteredColimits : PreservesFilteredColimits (forget SemiRingCat.{u}) := + Limits.comp_preservesFilteredColimits (forget₂ SemiRingCat MonCat) (forget MonCat.{u}) end @@ -197,17 +197,18 @@ def colimitCoconeIsColimit : IsColimit <| colimitCocone.{v, u} F where (Types.TypeMax.colimitCoconeIsColimit (F ⋙ forget CommSemiRingCat)).uniq ((forget CommSemiRingCat).mapCocone t) _ fun j => funext fun x => RingHom.congr_fun (h j) x -instance forget₂SemiRingPreservesFilteredColimits : +instance forget₂SemiRing_preservesFilteredColimits : PreservesFilteredColimits (forget₂ CommSemiRingCat SemiRingCat.{u}) where preserves_filtered_colimits {J hJ1 _} := letI : Category J := hJ1 { preservesColimit := fun {F} => - preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit.{u, u} F) + preservesColimit_of_preserves_colimit_cocone (colimitCoconeIsColimit.{u, u} F) (SemiRingCat.FilteredColimits.colimitCoconeIsColimit (F ⋙ forget₂ CommSemiRingCat SemiRingCat.{u})) } -instance forgetPreservesFilteredColimits : PreservesFilteredColimits (forget CommSemiRingCat.{u}) := - Limits.compPreservesFilteredColimits (forget₂ CommSemiRingCat SemiRingCat) +instance forget_preservesFilteredColimits : + PreservesFilteredColimits (forget CommSemiRingCat.{u}) := + Limits.comp_preservesFilteredColimits (forget₂ CommSemiRingCat SemiRingCat) (forget SemiRingCat.{u}) end @@ -259,17 +260,17 @@ def colimitCoconeIsColimit : IsColimit <| colimitCocone.{v, u} F where (Types.TypeMax.colimitCoconeIsColimit (F ⋙ forget RingCat)).uniq ((forget RingCat).mapCocone t) _ fun j => funext fun x => RingHom.congr_fun (h j) x -instance forget₂SemiRingPreservesFilteredColimits : +instance forget₂SemiRing_preservesFilteredColimits : PreservesFilteredColimits (forget₂ RingCat SemiRingCat.{u}) where preserves_filtered_colimits {J hJ1 _} := letI : Category J := hJ1 { preservesColimit := fun {F} => - preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit.{u, u} F) + preservesColimit_of_preserves_colimit_cocone (colimitCoconeIsColimit.{u, u} F) (SemiRingCat.FilteredColimits.colimitCoconeIsColimit (F ⋙ forget₂ RingCat SemiRingCat.{u})) } -instance forgetPreservesFilteredColimits : PreservesFilteredColimits (forget RingCat.{u}) := - Limits.compPreservesFilteredColimits (forget₂ RingCat SemiRingCat) (forget SemiRingCat.{u}) +instance forget_preservesFilteredColimits : PreservesFilteredColimits (forget RingCat.{u}) := + Limits.comp_preservesFilteredColimits (forget₂ RingCat SemiRingCat) (forget SemiRingCat.{u}) end @@ -319,16 +320,16 @@ def colimitCoconeIsColimit : IsColimit <| colimitCocone.{v, u} F where (Types.TypeMax.colimitCoconeIsColimit (F ⋙ forget CommRingCat)).uniq ((forget CommRingCat).mapCocone t) _ fun j => funext fun x => RingHom.congr_fun (h j) x -instance forget₂RingPreservesFilteredColimits : +instance forget₂Ring_preservesFilteredColimits : PreservesFilteredColimits (forget₂ CommRingCat RingCat.{u}) where preserves_filtered_colimits {J hJ1 _} := letI : Category J := hJ1 { preservesColimit := fun {F} => - preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit.{u, u} F) + preservesColimit_of_preserves_colimit_cocone (colimitCoconeIsColimit.{u, u} F) (RingCat.FilteredColimits.colimitCoconeIsColimit (F ⋙ forget₂ CommRingCat RingCat.{u})) } -instance forgetPreservesFilteredColimits : PreservesFilteredColimits (forget CommRingCat.{u}) := - Limits.compPreservesFilteredColimits (forget₂ CommRingCat RingCat) (forget RingCat.{u}) +instance forget_preservesFilteredColimits : PreservesFilteredColimits (forget CommRingCat.{u}) := + Limits.comp_preservesFilteredColimits (forget₂ CommRingCat RingCat) (forget RingCat.{u}) end diff --git a/Mathlib/Algebra/Category/Ring/Limits.lean b/Mathlib/Algebra/Category/Ring/Limits.lean index 13fea077bb9d5..7a46166affcf1 100644 --- a/Mathlib/Algebra/Category/Ring/Limits.lean +++ b/Mathlib/Algebra/Category/Ring/Limits.lean @@ -139,16 +139,16 @@ def forget₂AddCommMonPreservesLimitsAux : /-- The forgetful functor from semirings to additive commutative monoids preserves all limits. -/ -instance forget₂AddCommMonPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget₂AddCommMon_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget₂ SemiRingCat AddCommMonCat.{u}) where preservesLimitsOfShape {_ _} := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit.{v, u} F) (forget₂AddCommMonPreservesLimitsAux F) } -instance forget₂AddCommMonPreservesLimits : +instance forget₂AddCommMon_preservesLimits : PreservesLimits (forget₂ SemiRingCat AddCommMonCat.{u}) := - SemiRingCat.forget₂AddCommMonPreservesLimitsOfSize.{u, u} + SemiRingCat.forget₂AddCommMon_preservesLimitsOfSize.{u, u} /-- An auxiliary declaration to speed up typechecking. -/ @@ -160,27 +160,27 @@ def forget₂MonPreservesLimitsAux : /-- The forgetful functor from semirings to monoids preserves all limits. -/ -instance forget₂MonPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget₂Mon_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget₂ SemiRingCat MonCat.{u}) where preservesLimitsOfShape {_ _} := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone (limitConeIsLimit F) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit F) (forget₂MonPreservesLimitsAux.{v, u} F) } -instance forget₂MonPreservesLimits : PreservesLimits (forget₂ SemiRingCat MonCat.{u}) := - SemiRingCat.forget₂MonPreservesLimitsOfSize.{u, u} +instance forget₂Mon_preservesLimits : PreservesLimits (forget₂ SemiRingCat MonCat.{u}) := + SemiRingCat.forget₂Mon_preservesLimitsOfSize.{u, u} /-- The forgetful functor from semirings to types preserves all limits. -/ -instance forgetPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget SemiRingCat.{u}) where preservesLimitsOfShape {_ _} := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone (limitConeIsLimit F) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit F) (Types.Small.limitConeIsLimit.{v, u} (F ⋙ forget _)) } -instance forgetPreservesLimits : PreservesLimits (forget SemiRingCat.{u}) := - SemiRingCat.forgetPreservesLimitsOfSize.{u, u} +instance forget_preservesLimits : PreservesLimits (forget SemiRingCat.{u}) := + SemiRingCat.forget_preservesLimitsOfSize.{u, u} end SemiRingCat @@ -265,29 +265,29 @@ instance hasLimits : HasLimits CommSemiRingCat.{u} := /-- The forgetful functor from rings to semirings preserves all limits. -/ -instance forget₂SemiRingPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget₂SemiRing_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget₂ CommSemiRingCat SemiRingCat.{u}) where preservesLimitsOfShape {_ _} := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit.{v, u} F) (SemiRingCat.HasLimits.limitConeIsLimit (F ⋙ forget₂ _ SemiRingCat)) } -instance forget₂SemiRingPreservesLimits : +instance forget₂SemiRing_preservesLimits : PreservesLimits (forget₂ CommSemiRingCat SemiRingCat.{u}) := - CommSemiRingCat.forget₂SemiRingPreservesLimitsOfSize.{u, u} + CommSemiRingCat.forget₂SemiRing_preservesLimitsOfSize.{u, u} /-- The forgetful functor from rings to types preserves all limits. (That is, the underlying types could have been computed instead as limits in the category of types.) -/ -instance forgetPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget CommSemiRingCat.{u}) where preservesLimitsOfShape {_ _} := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit.{v, u} F) (Types.Small.limitConeIsLimit.{v, u} _) } -instance forgetPreservesLimits : PreservesLimits (forget CommSemiRingCat.{u}) := - CommSemiRingCat.forgetPreservesLimitsOfSize.{u, u} +instance forget_preservesLimits : PreservesLimits (forget CommSemiRingCat.{u}) := + CommSemiRingCat.forget_preservesLimitsOfSize.{u, u} end CommSemiRingCat @@ -374,15 +374,15 @@ instance hasLimits : HasLimits RingCat.{u} := /-- The forgetful functor from rings to semirings preserves all limits. -/ -instance forget₂SemiRingPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget₂SemiRing_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget₂ RingCat SemiRingCat.{u}) where preservesLimitsOfShape {_ _} := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit.{v, u} F) (SemiRingCat.HasLimits.limitConeIsLimit.{v, u} _) } -instance forget₂SemiRingPreservesLimits : PreservesLimits (forget₂ RingCat SemiRingCat.{u}) := - RingCat.forget₂SemiRingPreservesLimitsOfSize.{u, u} +instance forget₂SemiRing_preservesLimits : PreservesLimits (forget₂ RingCat SemiRingCat.{u}) := + RingCat.forget₂SemiRing_preservesLimitsOfSize.{u, u} /-- An auxiliary declaration to speed up typechecking. -/ @@ -396,29 +396,29 @@ def forget₂AddCommGroupPreservesLimitsAux : /-- The forgetful functor from rings to additive commutative groups preserves all limits. -/ -instance forget₂AddCommGroupPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget₂AddCommGroup_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{v, v} (forget₂ RingCat.{u} AddCommGrp.{u}) where preservesLimitsOfShape {_ _} := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit.{v, u} F) (forget₂AddCommGroupPreservesLimitsAux F) } -instance forget₂AddCommGroupPreservesLimits : +instance forget₂AddCommGroup_preservesLimits : PreservesLimits (forget₂ RingCat AddCommGrp.{u}) := - RingCat.forget₂AddCommGroupPreservesLimitsOfSize.{u, u} + RingCat.forget₂AddCommGroup_preservesLimitsOfSize.{u, u} /-- The forgetful functor from rings to types preserves all limits. (That is, the underlying types could have been computed instead as limits in the category of types.) -/ -instance forgetPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{v, v} (forget RingCat.{u}) where preservesLimitsOfShape {_ _} := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit.{v, u} F) (Types.Small.limitConeIsLimit.{v, u} _) } -instance forgetPreservesLimits : PreservesLimits (forget RingCat.{u}) := - RingCat.forgetPreservesLimitsOfSize.{u, u} +instance forget_preservesLimits : PreservesLimits (forget RingCat.{u}) := + RingCat.forget_preservesLimitsOfSize.{u, u} end RingCat @@ -514,15 +514,15 @@ instance hasLimits : HasLimits CommRingCat.{u} := /-- The forgetful functor from commutative rings to rings preserves all limits. (That is, the underlying rings could have been computed instead as limits in the category of rings.) -/ -instance forget₂RingPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget₂Ring_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget₂ CommRingCat RingCat.{u}) where preservesLimitsOfShape {_ _} := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone.{w, v} (limitConeIsLimit.{v, u} F) + preservesLimit_of_preserves_limit_cone.{w, v} (limitConeIsLimit.{v, u} F) (RingCat.limitConeIsLimit.{v, u} _) } -instance forget₂RingPreservesLimits : PreservesLimits (forget₂ CommRingCat RingCat.{u}) := - CommRingCat.forget₂RingPreservesLimitsOfSize.{u, u} +instance forget₂Ring_preservesLimits : PreservesLimits (forget₂ CommRingCat RingCat.{u}) := + CommRingCat.forget₂Ring_preservesLimitsOfSize.{u, u} /-- An auxiliary declaration to speed up typechecking. -/ @@ -536,28 +536,28 @@ def forget₂CommSemiRingPreservesLimitsAux : (That is, the underlying commutative semirings could have been computed instead as limits in the category of commutative semirings.) -/ -instance forget₂CommSemiRingPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget₂CommSemiRing_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget₂ CommRingCat CommSemiRingCat.{u}) where preservesLimitsOfShape {_ _} := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit.{v, u} F) (forget₂CommSemiRingPreservesLimitsAux.{v, u} F) } -instance forget₂CommSemiRingPreservesLimits : +instance forget₂CommSemiRing_preservesLimits : PreservesLimits (forget₂ CommRingCat CommSemiRingCat.{u}) := - CommRingCat.forget₂CommSemiRingPreservesLimitsOfSize.{u, u} + CommRingCat.forget₂CommSemiRing_preservesLimitsOfSize.{u, u} /-- The forgetful functor from commutative rings to types preserves all limits. (That is, the underlying types could have been computed instead as limits in the category of types.) -/ -instance forgetPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget CommRingCat.{u}) where preservesLimitsOfShape {_ _} := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone.{w, v} (limitConeIsLimit.{v, u} F) + preservesLimit_of_preserves_limit_cone.{w, v} (limitConeIsLimit.{v, u} F) (Types.Small.limitConeIsLimit.{v, u} _) } -instance forgetPreservesLimits : PreservesLimits (forget CommRingCat.{u}) := - CommRingCat.forgetPreservesLimitsOfSize.{u, u} +instance forget_preservesLimits : PreservesLimits (forget CommRingCat.{u}) := + CommRingCat.forget_preservesLimitsOfSize.{u, u} end CommRingCat diff --git a/Mathlib/Algebra/Homology/DerivedCategory/Ext/Basic.lean b/Mathlib/Algebra/Homology/DerivedCategory/Ext/Basic.lean index 774df032bb48c..b5130066b0f40 100644 --- a/Mathlib/Algebra/Homology/DerivedCategory/Ext/Basic.lean +++ b/Mathlib/Algebra/Homology/DerivedCategory/Ext/Basic.lean @@ -251,7 +251,7 @@ lemma mk₀_zero : mk₀ (0 : X ⟶ Y) = 0 := by section -attribute [local instance] preservesBinaryBiproductsOfPreservesBiproducts in +attribute [local instance] preservesBinaryBiproducts_of_preservesBiproducts in lemma biprod_ext {X₁ X₂ : C} {α β : Ext (X₁ ⊞ X₂) Y n} (h₁ : (mk₀ biprod.inl).comp α (zero_add n) = (mk₀ biprod.inl).comp β (zero_add n)) (h₂ : (mk₀ biprod.inr).comp α (zero_add n) = (mk₀ biprod.inr).comp β (zero_add n)) : diff --git a/Mathlib/Algebra/Homology/HomologicalComplexBiprod.lean b/Mathlib/Algebra/Homology/HomologicalComplexBiprod.lean index cb8b60d15b0e0..48e36902658d6 100644 --- a/Mathlib/Algebra/Homology/HomologicalComplexBiprod.lean +++ b/Mathlib/Algebra/Homology/HomologicalComplexBiprod.lean @@ -36,7 +36,7 @@ instance (i : ι) : HasColimit ((pair K L) ⋙ (eval C c i)) := by instance : HasBinaryBiproduct K L := HasBinaryBiproduct.of_hasBinaryProduct _ _ instance (i : ι) : PreservesBinaryBiproduct K L (eval C c i) := - preservesBinaryBiproductOfPreservesBinaryProduct _ + preservesBinaryBiproduct_of_preservesBinaryProduct _ /-- The canonical isomorphism `(K ⊞ L).X i ≅ (K.X i) ⊞ (L.X i)`. -/ noncomputable def biprodXIso (i : ι) : (K ⊞ L).X i ≅ (K.X i) ⊞ (L.X i) := diff --git a/Mathlib/Algebra/Homology/HomologicalComplexLimits.lean b/Mathlib/Algebra/Homology/HomologicalComplexLimits.lean index 0e2a446e27264..078140361610a 100644 --- a/Mathlib/Algebra/Homology/HomologicalComplexLimits.lean +++ b/Mathlib/Algebra/Homology/HomologicalComplexLimits.lean @@ -79,7 +79,7 @@ noncomputable def isLimitConeOfHasLimitEval : IsLimit (coneOfHasLimitEval F) := instance : HasLimit F := ⟨⟨⟨_, isLimitConeOfHasLimitEval F⟩⟩⟩ noncomputable instance (n : ι) : PreservesLimit F (eval C c n) := - preservesLimitOfPreservesLimitCone (isLimitConeOfHasLimitEval F) (limit.isLimit _) + preservesLimit_of_preserves_limit_cone (isLimitConeOfHasLimitEval F) (limit.isLimit _) end @@ -155,7 +155,7 @@ noncomputable def isColimitCoconeOfHasColimitEval : IsColimit (coconeOfHasColimi instance : HasColimit F := ⟨⟨⟨_, isColimitCoconeOfHasColimitEval F⟩⟩⟩ noncomputable instance (n : ι) : PreservesColimit F (eval C c n) := - preservesColimitOfPreservesColimitCocone (isColimitCoconeOfHasColimitEval F) + preservesColimit_of_preserves_colimit_cocone (isColimitCoconeOfHasColimitEval F) (colimit.isColimit _) end @@ -178,39 +178,39 @@ instance [HasFiniteColimits C] {K L : HomologicalComplex C c} (φ : K ⟶ L) [Ep /-- A functor `D ⥤ HomologicalComplex C c` preserves limits of shape `J` if for any `i`, `G ⋙ eval C c i` does. -/ -def preservesLimitsOfShapeOfEval {D : Type*} [Category D] +lemma preservesLimitsOfShape_of_eval {D : Type*} [Category D] (G : D ⥤ HomologicalComplex C c) (_ : ∀ (i : ι), PreservesLimitsOfShape J (G ⋙ eval C c i)) : PreservesLimitsOfShape J G := - ⟨fun {_} => ⟨fun hs ↦ isLimitOfEval _ _ - (fun i => isLimitOfPreserves (G ⋙ eval C c i) hs)⟩⟩ + ⟨fun {_} => ⟨fun hs ↦ ⟨isLimitOfEval _ _ + (fun i => isLimitOfPreserves (G ⋙ eval C c i) hs)⟩⟩⟩ /-- A functor `D ⥤ HomologicalComplex C c` preserves colimits of shape `J` if for any `i`, `G ⋙ eval C c i` does. -/ -def preservesColimitsOfShapeOfEval {D : Type*} [Category D] +lemma preservesColimitsOfShape_of_eval {D : Type*} [Category D] (G : D ⥤ HomologicalComplex C c) (_ : ∀ (i : ι), PreservesColimitsOfShape J (G ⋙ eval C c i)) : PreservesColimitsOfShape J G := - ⟨fun {_} => ⟨fun hs ↦ isColimitOfEval _ _ - (fun i => isColimitOfPreserves (G ⋙ eval C c i) hs)⟩⟩ + ⟨fun {_} => ⟨fun hs ↦ ⟨isColimitOfEval _ _ + (fun i => isColimitOfPreserves (G ⋙ eval C c i) hs)⟩⟩⟩ section variable [HasZeroObject C] [DecidableEq ι] (i : ι) noncomputable instance : PreservesLimitsOfShape J (single C c i) := - preservesLimitsOfShapeOfEval _ (fun j => by + preservesLimitsOfShape_of_eval _ (fun j => by by_cases h : j = i · subst h - exact preservesLimitsOfShapeOfNatIso (singleCompEvalIsoSelf C c j).symm - · exact Functor.preservesLimitsOfShapeOfIsZero _ (isZero_single_comp_eval C c _ _ h) _) + exact preservesLimitsOfShape_of_natIso (singleCompEvalIsoSelf C c j).symm + · exact Functor.preservesLimitsOfShape_of_isZero _ (isZero_single_comp_eval C c _ _ h) _) noncomputable instance : PreservesColimitsOfShape J (single C c i) := - preservesColimitsOfShapeOfEval _ (fun j => by + preservesColimitsOfShape_of_eval _ (fun j => by by_cases h : j = i · subst h - exact preservesColimitsOfShapeOfNatIso (singleCompEvalIsoSelf C c j).symm - · exact Functor.preservesColimitsOfShapeOfIsZero _ (isZero_single_comp_eval C c _ _ h) _) + exact preservesColimitsOfShape_of_natIso (singleCompEvalIsoSelf C c j).symm + · exact Functor.preservesColimitsOfShape_of_isZero _ (isZero_single_comp_eval C c _ _ h) _) noncomputable instance : PreservesFiniteLimits (single C c i) := ⟨by intros; infer_instance⟩ diff --git a/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean b/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean index 51fd7a82b8752..dfd6dc033079a 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean @@ -139,10 +139,10 @@ variable (D : SnakeInput C) lemma δ_apply (x₃ : D.L₀.X₃) (x₂ : D.L₁.X₂) (x₁ : D.L₂.X₁) (h₂ : D.L₁.g x₂ = D.v₀₁.τ₃ x₃) (h₁ : D.L₂.f x₁ = D.v₁₂.τ₂ x₂) : D.δ x₃ = D.v₂₃.τ₁ x₁ := by - have := (forget₂ C Ab).preservesFiniteLimitsOfPreservesHomology + have := (forget₂ C Ab).preservesFiniteLimits_of_preservesHomology have : PreservesFiniteLimits (forget C) := by have : forget₂ C Ab ⋙ forget Ab = forget C := HasForget₂.forget_comp - simpa only [← this] using compPreservesFiniteLimits _ _ + simpa only [← this] using comp_preservesFiniteLimits _ _ have eq := congr_fun ((forget C).congr_map D.snd_δ) (Limits.Concrete.pullbackMk D.L₁.g D.v₀₁.τ₃ x₂ x₃ h₂) have eq₁ := Concrete.pullbackMk_fst D.L₁.g D.v₀₁.τ₃ x₂ x₃ h₂ diff --git a/Mathlib/Algebra/Homology/ShortComplex/ExactFunctor.lean b/Mathlib/Algebra/Homology/ShortComplex/ExactFunctor.lean index d61079ee62930..06141d6538e48 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/ExactFunctor.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/ExactFunctor.lean @@ -19,9 +19,9 @@ also preserves finite limits and finite colimits. Let `F : C ⥤ D` be an additive functor: -- `Functor.preservesFiniteLimitsOfPreservesHomology`: if `F` preserves homology, then `F` preserves - finite limits. -- `Functor.preservesFiniteColimitsOfPreservesHomology`: if `F` preserves homology, then `F` +- `Functor.preservesFiniteLimits_of_preservesHomology`: if `F` preserves homology, + then `F` preserves finite limits. +- `Functor.preservesFiniteColimits_of_preservesHomology`: if `F` preserves homology, then `F` preserves finite colimits. If we further assume that `C` and `D` are abelian categories, then we have: @@ -63,24 +63,24 @@ variable {C D : Type*} [Category C] [Category D] [Preadditive C] [Preadditive D] (F : C ⥤ D) [F.Additive] [F.PreservesHomology] [HasZeroObject C] /-- An additive functor which preserves homology preserves finite limits. -/ -noncomputable def preservesFiniteLimitsOfPreservesHomology +lemma preservesFiniteLimits_of_preservesHomology [HasFiniteProducts C] [HasKernels C] : PreservesFiniteLimits F := by have := fun {X Y : C} (f : X ⟶ Y) ↦ PreservesHomology.preservesKernel F f have : HasBinaryBiproducts C := HasBinaryBiproducts.of_hasBinaryProducts have : HasEqualizers C := Preadditive.hasEqualizers_of_hasKernels have : HasZeroObject D := ⟨F.obj 0, by rw [IsZero.iff_id_eq_zero, ← F.map_id, id_zero, F.map_zero]⟩ - exact preservesFiniteLimitsOfPreservesKernels F + exact preservesFiniteLimits_of_preservesKernels F /-- An additive which preserves homology preserves finite colimits. -/ -noncomputable def preservesFiniteColimitsOfPreservesHomology +lemma preservesFiniteColimits_of_preservesHomology [HasFiniteCoproducts C] [HasCokernels C] : PreservesFiniteColimits F := by have := fun {X Y : C} (f : X ⟶ Y) ↦ PreservesHomology.preservesCokernel F f have : HasBinaryBiproducts C := HasBinaryBiproducts.of_hasBinaryCoproducts have : HasCoequalizers C := Preadditive.hasCoequalizers_of_hasCokernels have : HasZeroObject D := ⟨F.obj 0, by rw [IsZero.iff_id_eq_zero, ← F.map_id, id_zero, F.map_zero]⟩ - exact preservesFiniteColimitsOfPreservesCokernels F + exact preservesFiniteColimits_of_preservesCokernels F end @@ -112,8 +112,8 @@ lemma preservesFiniteLimits_tfae : List.TFAE [ ∀ (S : ShortComplex C), S.ShortExact → (S.map F).Exact ∧ Mono (F.map S.f), ∀ (S : ShortComplex C), S.Exact ∧ Mono S.f → (S.map F).Exact ∧ Mono (F.map S.f), - ∀ ⦃X Y : C⦄ (f : X ⟶ Y), Nonempty <| PreservesLimit (parallelPair f 0) F, - Nonempty <| PreservesFiniteLimits F + ∀ ⦃X Y : C⦄ (f : X ⟶ Y), PreservesLimit (parallelPair f 0) F, + PreservesFiniteLimits F ] := by tfae_have 1 → 2 | hF, S, ⟨hS, hf⟩ => by @@ -130,7 +130,7 @@ lemma preservesFiniteLimits_tfae : List.TFAE tfae_have 2 → 3 | hF, X, Y, f => by - refine ⟨preservesLimitOfPreservesLimitCone (kernelIsKernel f) ?_⟩ + refine preservesLimit_of_preserves_limit_cone (kernelIsKernel f) ?_ apply (KernelFork.isLimitMapConeEquiv _ F).2 let S := ShortComplex.mk _ _ (kernel.condition f) let hS := hF S ⟨exact_kernel f, inferInstance⟩ @@ -139,8 +139,7 @@ lemma preservesFiniteLimits_tfae : List.TFAE tfae_have 3 → 4 | hF => by - have := fun X Y (f : X ⟶ Y) ↦ (hF f).some - exact ⟨preservesFiniteLimitsOfPreservesKernels F⟩ + exact preservesFiniteLimits_of_preservesKernels F tfae_have 4 → 1 | ⟨_⟩, S, hS => @@ -171,8 +170,8 @@ lemma preservesFiniteColimits_tfae : List.TFAE [ ∀ (S : ShortComplex C), S.ShortExact → (S.map F).Exact ∧ Epi (F.map S.g), ∀ (S : ShortComplex C), S.Exact ∧ Epi S.g → (S.map F).Exact ∧ Epi (F.map S.g), - ∀ ⦃X Y : C⦄ (f : X ⟶ Y), Nonempty <| PreservesColimit (parallelPair f 0) F, - Nonempty <| PreservesFiniteColimits F + ∀ ⦃X Y : C⦄ (f : X ⟶ Y), PreservesColimit (parallelPair f 0) F, + PreservesFiniteColimits F ] := by tfae_have 1 → 2 | hF, S, ⟨hS, hf⟩ => by @@ -189,7 +188,7 @@ lemma preservesFiniteColimits_tfae : List.TFAE tfae_have 2 → 3 | hF, X, Y, f => by - refine ⟨preservesColimitOfPreservesColimitCocone (cokernelIsCokernel f) ?_⟩ + refine preservesColimit_of_preserves_colimit_cocone (cokernelIsCokernel f) ?_ apply (CokernelCofork.isColimitMapCoconeEquiv _ F).2 let S := ShortComplex.mk _ _ (cokernel.condition f) let hS := hF S ⟨exact_cokernel f, inferInstance⟩ @@ -198,8 +197,7 @@ lemma preservesFiniteColimits_tfae : List.TFAE tfae_have 3 → 4 | hF => by - have := fun X Y (f : X ⟶ Y) ↦ (hF f).some - exact ⟨preservesFiniteColimitsOfPreservesCokernels F⟩ + exact preservesFiniteColimits_of_preservesCokernels F tfae_have 4 → 1 | ⟨_⟩, S, hS => (S.map F).exact_and_epi_g_iff_g_is_cokernel |>.2 @@ -219,18 +217,18 @@ lemma exact_tfae : List.TFAE [ ∀ (S : ShortComplex C), S.ShortExact → (S.map F).ShortExact, ∀ (S : ShortComplex C), S.Exact → (S.map F).Exact, - Nonempty (PreservesHomology F), - Nonempty (PreservesFiniteLimits F) ∧ Nonempty (PreservesFiniteColimits F) + PreservesHomology F, + PreservesFiniteLimits F ∧ PreservesFiniteColimits F ] := by tfae_have 1 → 3 | hF => by refine ⟨fun {X Y} f ↦ ?_, fun {X Y} f ↦ ?_⟩ · have h := (preservesFiniteLimits_tfae F |>.out 0 2 |>.1 fun S hS ↦ And.intro (hF S hS).exact (hF S hS).mono_f) - exact h f |>.some + exact h f · have h := (preservesFiniteColimits_tfae F |>.out 0 2 |>.1 fun S hS ↦ And.intro (hF S hS).exact (hF S hS).epi_g) - exact h f |>.some + exact h f tfae_have 2 → 1 | hF, S, hS => by @@ -241,11 +239,11 @@ lemma exact_tfae : List.TFAE exact ⟨hF S hS.exact⟩ tfae_have 3 → 4 - | ⟨h⟩ => ⟨⟨preservesFiniteLimitsOfPreservesHomology F⟩, - ⟨preservesFiniteColimitsOfPreservesHomology F⟩⟩ + | h => ⟨preservesFiniteLimits_of_preservesHomology F, + preservesFiniteColimits_of_preservesHomology F⟩ tfae_have 4 → 2 - | ⟨⟨h1⟩, ⟨h2⟩⟩, _, h => h.map F + | ⟨h1, h2⟩, _, h => h.map F tfae_finish diff --git a/Mathlib/Algebra/Homology/ShortComplex/Limits.lean b/Mathlib/Algebra/Homology/ShortComplex/Limits.lean index ba94784012560..1f94e0fc71c5e 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/Limits.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/Limits.lean @@ -14,8 +14,6 @@ import Mathlib.CategoryTheory.Limits.Preserves.Finite In this file, it is shown if a category `C` with zero morphisms has limits of a certain shape `J`, then it is also the case of the category `ShortComplex C`. -TODO (@rioujoel): Do the same for colimits. - -/ namespace CategoryTheory @@ -95,13 +93,13 @@ noncomputable def isLimitLimitCone : IsLimit (limitCone F) := instance hasLimit_of_hasLimitπ : HasLimit F := ⟨⟨⟨_, isLimitLimitCone _⟩⟩⟩ noncomputable instance : PreservesLimit F π₁ := - preservesLimitOfPreservesLimitCone (isLimitLimitCone F) (isLimitπ₁MapConeLimitCone F) + preservesLimit_of_preserves_limit_cone (isLimitLimitCone F) (isLimitπ₁MapConeLimitCone F) noncomputable instance : PreservesLimit F π₂ := - preservesLimitOfPreservesLimitCone (isLimitLimitCone F) (isLimitπ₂MapConeLimitCone F) + preservesLimit_of_preserves_limit_cone (isLimitLimitCone F) (isLimitπ₂MapConeLimitCone F) noncomputable instance : PreservesLimit F π₃ := - preservesLimitOfPreservesLimitCone (isLimitLimitCone F) (isLimitπ₃MapConeLimitCone F) + preservesLimit_of_preserves_limit_cone (isLimitLimitCone F) (isLimitπ₃MapConeLimitCone F) end @@ -231,15 +229,15 @@ noncomputable def isColimitColimitCocone : IsColimit (colimitCocone F) := instance hasColimit_of_hasColimitπ : HasColimit F := ⟨⟨⟨_, isColimitColimitCocone _⟩⟩⟩ noncomputable instance : PreservesColimit F π₁ := - preservesColimitOfPreservesColimitCocone (isColimitColimitCocone F) + preservesColimit_of_preserves_colimit_cocone (isColimitColimitCocone F) (isColimitπ₁MapCoconeColimitCocone F) noncomputable instance : PreservesColimit F π₂ := - preservesColimitOfPreservesColimitCocone (isColimitColimitCocone F) + preservesColimit_of_preserves_colimit_cocone (isColimitColimitCocone F) (isColimitπ₂MapCoconeColimitCocone F) noncomputable instance : PreservesColimit F π₃ := - preservesColimitOfPreservesColimitCocone (isColimitColimitCocone F) + preservesColimit_of_preserves_colimit_cocone (isColimitColimitCocone F) (isColimitπ₃MapCoconeColimitCocone F) end diff --git a/Mathlib/Algebra/Homology/ShortComplex/PreservesHomology.lean b/Mathlib/Algebra/Homology/ShortComplex/PreservesHomology.lean index 5f870fcf8b0c0..eacf12a85c953 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/PreservesHomology.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/PreservesHomology.lean @@ -34,7 +34,7 @@ namespace Functor variable (F : C ⥤ D) /-- A functor preserves homology when it preserves both kernels and cokernels. -/ -class PreservesHomology (F : C ⥤ D) [PreservesZeroMorphisms F] where +class PreservesHomology (F : C ⥤ D) [PreservesZeroMorphisms F] : Prop where /-- the functor preserves kernels -/ preservesKernels ⦃X Y : C⦄ (f : X ⟶ Y) : PreservesLimit (parallelPair f 0) F := by infer_instance @@ -45,12 +45,12 @@ class PreservesHomology (F : C ⥤ D) [PreservesZeroMorphisms F] where variable [PreservesZeroMorphisms F] /-- A functor which preserves homology preserves kernels. -/ -def PreservesHomology.preservesKernel [F.PreservesHomology] {X Y : C} (f : X ⟶ Y) : +lemma PreservesHomology.preservesKernel [F.PreservesHomology] {X Y : C} (f : X ⟶ Y) : PreservesLimit (parallelPair f 0) F := PreservesHomology.preservesKernels _ /-- A functor which preserves homology preserves cokernels. -/ -def PreservesHomology.preservesCokernel [F.PreservesHomology] {X Y : C} (f : X ⟶ Y) : +lemma PreservesHomology.preservesCokernel [F.PreservesHomology] {X Y : C} (f : X ⟶ Y) : PreservesColimit (parallelPair f 0) F := PreservesHomology.preservesCokernels _ @@ -70,7 +70,7 @@ variable (h : S.LeftHomologyData) (F : C ⥤ D) /-- A left homology data `h` of a short complex `S` is preserved by a functor `F` is `F` preserves the kernel of `S.g : S.X₂ ⟶ S.X₃` and the cokernel of `h.f' : S.X₁ ⟶ h.K`. -/ -class IsPreservedBy [F.PreservesZeroMorphisms] where +class IsPreservedBy [F.PreservesZeroMorphisms] : Prop where /-- the functor preserves the kernel of `S.g : S.X₂ ⟶ S.X₃`. -/ g : PreservesLimit (parallelPair S.g 0) F /-- the functor preserves the cokernel of `h.f' : S.X₁ ⟶ h.K`. -/ @@ -78,21 +78,22 @@ class IsPreservedBy [F.PreservesZeroMorphisms] where variable [F.PreservesZeroMorphisms] -noncomputable instance isPreservedByOfPreservesHomology [F.PreservesHomology] : +noncomputable instance isPreservedBy_of_preservesHomology [F.PreservesHomology] : h.IsPreservedBy F where g := Functor.PreservesHomology.preservesKernel _ _ f' := Functor.PreservesHomology.preservesCokernel _ _ variable [h.IsPreservedBy F] +include h in /-- When a left homology data is preserved by a functor `F`, this functor preserves the kernel of `S.g : S.X₂ ⟶ S.X₃`. -/ -def IsPreservedBy.hg : PreservesLimit (parallelPair S.g 0) F := +lemma IsPreservedBy.hg : PreservesLimit (parallelPair S.g 0) F := @IsPreservedBy.g _ _ _ _ _ _ _ h F _ _ /-- When a left homology data `h` is preserved by a functor `F`, this functor preserves the cokernel of `h.f' : S.X₁ ⟶ h.K`. -/ -def IsPreservedBy.hf' : PreservesColimit (parallelPair h.f' 0) F := IsPreservedBy.f' +lemma IsPreservedBy.hf' : PreservesColimit (parallelPair h.f' 0) F := IsPreservedBy.f' /-- When a left homology data `h` of a short complex `S` is preserved by a functor `F`, this is the induced left homology data `h.map F` for the short complex `S.map F`. -/ @@ -149,7 +150,7 @@ variable (h : S.RightHomologyData) (F : C ⥤ D) /-- A right homology data `h` of a short complex `S` is preserved by a functor `F` is `F` preserves the cokernel of `S.f : S.X₁ ⟶ S.X₂` and the kernel of `h.g' : h.Q ⟶ S.X₃`. -/ -class IsPreservedBy [F.PreservesZeroMorphisms] where +class IsPreservedBy [F.PreservesZeroMorphisms] : Prop where /-- the functor preserves the cokernel of `S.f : S.X₁ ⟶ S.X₂`. -/ f : PreservesColimit (parallelPair S.f 0) F /-- the functor preserves the kernel of `h.g' : h.Q ⟶ S.X₃`. -/ @@ -157,21 +158,22 @@ class IsPreservedBy [F.PreservesZeroMorphisms] where variable [F.PreservesZeroMorphisms] -noncomputable instance isPreservedByOfPreservesHomology [F.PreservesHomology] : +noncomputable instance isPreservedBy_of_preservesHomology [F.PreservesHomology] : h.IsPreservedBy F where f := Functor.PreservesHomology.preservesCokernel F _ g' := Functor.PreservesHomology.preservesKernel F _ variable [h.IsPreservedBy F] +include h in /-- When a right homology data is preserved by a functor `F`, this functor preserves the cokernel of `S.f : S.X₁ ⟶ S.X₂`. -/ -def IsPreservedBy.hf : PreservesColimit (parallelPair S.f 0) F := +lemma IsPreservedBy.hf : PreservesColimit (parallelPair S.f 0) F := @IsPreservedBy.f _ _ _ _ _ _ _ h F _ _ /-- When a right homology data `h` is preserved by a functor `F`, this functor preserves the kernel of `h.g' : h.Q ⟶ S.X₃`. -/ -def IsPreservedBy.hg' : PreservesLimit (parallelPair h.g' 0) F := +lemma IsPreservedBy.hg' : PreservesLimit (parallelPair h.g' 0) F := @IsPreservedBy.g' _ _ _ _ _ _ _ h F _ _ /-- When a right homology data `h` of a short complex `S` is preserved by a functor `F`, @@ -256,27 +258,27 @@ variable (F : C ⥤ D) [PreservesZeroMorphisms F] (S : ShortComplex C) {S₁ S /-- A functor preserves the left homology of a short complex `S` if it preserves all the left homology data of `S`. -/ -class PreservesLeftHomologyOf where +class PreservesLeftHomologyOf : Prop where /-- the functor preserves all the left homology data of the short complex -/ isPreservedBy : ∀ (h : S.LeftHomologyData), h.IsPreservedBy F /-- A functor preserves the right homology of a short complex `S` if it preserves all the right homology data of `S`. -/ -class PreservesRightHomologyOf where +class PreservesRightHomologyOf : Prop where /-- the functor preserves all the right homology data of the short complex -/ isPreservedBy : ∀ (h : S.RightHomologyData), h.IsPreservedBy F -noncomputable instance PreservesHomology.preservesLeftHomologyOf [F.PreservesHomology] : +instance PreservesHomology.preservesLeftHomologyOf [F.PreservesHomology] : F.PreservesLeftHomologyOf S := ⟨inferInstance⟩ -noncomputable instance PreservesHomology.preservesRightHomologyOf [F.PreservesHomology] : +instance PreservesHomology.preservesRightHomologyOf [F.PreservesHomology] : F.PreservesRightHomologyOf S := ⟨inferInstance⟩ variable {S} /-- If a functor preserves a certain left homology data of a short complex `S`, then it preserves the left homology of `S`. -/ -def PreservesLeftHomologyOf.mk' (h : S.LeftHomologyData) [h.IsPreservedBy F] : +lemma PreservesLeftHomologyOf.mk' (h : S.LeftHomologyData) [h.IsPreservedBy F] : F.PreservesLeftHomologyOf S where isPreservedBy h' := { g := ShortComplex.LeftHomologyData.IsPreservedBy.hg h F @@ -285,11 +287,11 @@ def PreservesLeftHomologyOf.mk' (h : S.LeftHomologyData) [h.IsPreservedBy F] : let e : parallelPair h.f' 0 ≅ parallelPair h'.f' 0 := parallelPair.ext (Iso.refl _) (ShortComplex.cyclesMapIso' (Iso.refl S) h h') (by simp) (by simp) - exact preservesColimitOfIsoDiagram F e } + exact preservesColimit_of_iso_diagram F e } /-- If a functor preserves a certain right homology data of a short complex `S`, then it preserves the right homology of `S`. -/ -def PreservesRightHomologyOf.mk' (h : S.RightHomologyData) [h.IsPreservedBy F] : +lemma PreservesRightHomologyOf.mk' (h : S.RightHomologyData) [h.IsPreservedBy F] : F.PreservesRightHomologyOf S where isPreservedBy h' := { f := ShortComplex.RightHomologyData.IsPreservedBy.hf h F @@ -298,7 +300,7 @@ def PreservesRightHomologyOf.mk' (h : S.RightHomologyData) [h.IsPreservedBy F] : let e : parallelPair h.g' 0 ≅ parallelPair h'.g' 0 := parallelPair.ext (ShortComplex.opcyclesMapIso' (Iso.refl S) h h') (Iso.refl _) (by simp) (by simp) - exact preservesLimitOfIsoDiagram F e } + exact preservesLimit_of_iso_diagram F e } end Functor @@ -307,11 +309,11 @@ namespace ShortComplex variable {S : ShortComplex C} (h₁ : S.LeftHomologyData) (h₂ : S.RightHomologyData) (F : C ⥤ D) [F.PreservesZeroMorphisms] -instance LeftHomologyData.isPreservedByOfPreserves [F.PreservesLeftHomologyOf S] : +instance LeftHomologyData.isPreservedBy_of_preserves [F.PreservesLeftHomologyOf S] : h₁.IsPreservedBy F := Functor.PreservesLeftHomologyOf.isPreservedBy _ -instance RightHomologyData.isPreservedByOfPreserves [F.PreservesRightHomologyOf S] : +instance RightHomologyData.isPreservedBy_of_preserves [F.PreservesRightHomologyOf S] : h₂.IsPreservedBy F := Functor.PreservesRightHomologyOf.isPreservedBy _ @@ -805,25 +807,25 @@ variable (F : C ⥤ D) [F.PreservesZeroMorphisms] (S : ShortComplex C) /-- If a short complex `S` is such that `S.f = 0` and that the kernel of `S.g` is preserved by a functor `F`, then `F` preserves the left homology of `S`. -/ -noncomputable def preservesLeftHomologyOfZerof (hf : S.f = 0) +lemma preservesLeftHomology_of_zero_f (hf : S.f = 0) [PreservesLimit (parallelPair S.g 0) F] : F.PreservesLeftHomologyOf S := ⟨fun h => { g := by infer_instance - f' := Limits.preservesCokernelZero' _ _ + f' := Limits.preservesCokernel_zero' _ _ (by rw [← cancel_mono h.i, h.f'_i, zero_comp, hf]) }⟩ /-- If a short complex `S` is such that `S.g = 0` and that the cokernel of `S.f` is preserved by a functor `F`, then `F` preserves the right homology of `S`. -/ -noncomputable def preservesRightHomologyOfZerog (hg : S.g = 0) +lemma preservesRightHomology_of_zero_g (hg : S.g = 0) [PreservesColimit (parallelPair S.f 0) F] : F.PreservesRightHomologyOf S := ⟨fun h => { f := by infer_instance - g' := Limits.preservesKernelZero' _ _ + g' := Limits.preservesKernel_zero' _ _ (by rw [← cancel_epi h.p, h.p_g', comp_zero, hg]) }⟩ /-- If a short complex `S` is such that `S.g = 0` and that the cokernel of `S.f` is preserved by a functor `F`, then `F` preserves the left homology of `S`. -/ -noncomputable def preservesLeftHomologyOfZerog (hg : S.g = 0) +lemma preservesLeftHomology_of_zero_g (hg : S.g = 0) [PreservesColimit (parallelPair S.f 0) F] : F.PreservesLeftHomologyOf S := ⟨fun h => { g := by @@ -833,11 +835,11 @@ noncomputable def preservesLeftHomologyOfZerog (hg : S.g = 0) have := h.isIso_i hg let e : parallelPair h.f' 0 ≅ parallelPair S.f 0 := parallelPair.ext (Iso.refl _) (asIso h.i) (by aesop_cat) (by aesop_cat) - exact Limits.preservesColimitOfIsoDiagram F e.symm}⟩ + exact Limits.preservesColimit_of_iso_diagram F e.symm}⟩ /-- If a short complex `S` is such that `S.f = 0` and that the kernel of `S.g` is preserved by a functor `F`, then `F` preserves the right homology of `S`. -/ -noncomputable def preservesRightHomologyOfZerof (hf : S.f = 0) +lemma preservesRightHomology_of_zero_f (hf : S.f = 0) [PreservesLimit (parallelPair S.g 0) F] : F.PreservesRightHomologyOf S := ⟨fun h => { f := by @@ -847,7 +849,7 @@ noncomputable def preservesRightHomologyOfZerof (hf : S.f = 0) have := h.isIso_p hf let e : parallelPair S.g 0 ≅ parallelPair h.g' 0 := parallelPair.ext (asIso h.p) (Iso.refl _) (by aesop_cat) (by aesop_cat) - exact Limits.preservesLimitOfIsoDiagram F e }⟩ + exact Limits.preservesLimit_of_iso_diagram F e }⟩ end Functor diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index 956a2db40bc22..a83addf8247e5 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -194,7 +194,7 @@ noncomputable instance Γ_preservesLimits : PreservesLimits Γ.{u}.rightOp := in noncomputable instance forgetToScheme_preservesLimits : PreservesLimits forgetToScheme := by apply (config := { allowSynthFailures := true }) - @preservesLimitsOfNatIso _ _ _ _ _ _ + @preservesLimits_of_natIso _ _ _ _ _ _ (isoWhiskerRight equivCommRingCat.unitIso forgetToScheme).symm change PreservesLimits (equivCommRingCat.functor ⋙ Scheme.Spec) infer_instance diff --git a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean index 3b5e437273378..6c8faa057ea86 100644 --- a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean +++ b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean @@ -483,10 +483,10 @@ alias ΓSpec.adjunction_unit_map_basicOpen := Scheme.toSpecΓ_preimage_basicOpen /-- Spec preserves limits. -/ instance : Limits.PreservesLimits Spec.toLocallyRingedSpace := - ΓSpec.locallyRingedSpaceAdjunction.rightAdjointPreservesLimits + ΓSpec.locallyRingedSpaceAdjunction.rightAdjoint_preservesLimits instance Spec.preservesLimits : Limits.PreservesLimits Scheme.Spec := - ΓSpec.adjunction.rightAdjointPreservesLimits + ΓSpec.adjunction.rightAdjoint_preservesLimits /-- The functor `Spec.toLocallyRingedSpace : CommRingCatᵒᵖ ⥤ LocallyRingedSpace` is fully faithful.-/ diff --git a/Mathlib/AlgebraicGeometry/Limits.lean b/Mathlib/AlgebraicGeometry/Limits.lean index d94a986f233f9..6c634b3aadaac 100644 --- a/Mathlib/AlgebraicGeometry/Limits.lean +++ b/Mathlib/AlgebraicGeometry/Limits.lean @@ -197,8 +197,7 @@ instance : CreatesColimitsOfShape (Discrete ι) Scheme.forgetToLocallyRingedSpac intro K exact createsColimitOfIsoDiagram _ (Discrete.natIsoFunctor (F := K)).symm -noncomputable -instance : PreservesColimitsOfShape (Discrete ι) Scheme.forgetToTop := +instance : PreservesColimitsOfShape (Discrete ι) Scheme.forgetToTop.{u} := inferInstanceAs (PreservesColimitsOfShape (Discrete ι) (Scheme.forgetToLocallyRingedSpace ⋙ LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forget CommRingCat)) @@ -208,12 +207,15 @@ instance : HasCoproducts.{u} Scheme.{u} := instance : HasCoproducts.{0} Scheme.{u} := has_smallest_coproducts_of_hasCoproducts noncomputable -instance {ι : Type} : PreservesColimitsOfShape (Discrete ι) Scheme.forgetToTop := - preservesColimitsOfShapeOfEquiv (Discrete.equivalence Equiv.ulift : Discrete (ULift.{u} ι) ≌ _) _ +instance {ι : Type} : PreservesColimitsOfShape (Discrete ι) Scheme.forgetToTop.{u} := + preservesColimitsOfShape_of_equiv + (Discrete.equivalence Equiv.ulift : Discrete (ULift.{u} ι) ≌ _) _ noncomputable -instance {ι : Type} : PreservesColimitsOfShape (Discrete ι) Scheme.forgetToLocallyRingedSpace := - preservesColimitsOfShapeOfEquiv (Discrete.equivalence Equiv.ulift : Discrete (ULift.{u} ι) ≌ _) _ +instance {ι : Type} : + PreservesColimitsOfShape (Discrete ι) Scheme.forgetToLocallyRingedSpace.{u} := + preservesColimitsOfShape_of_equiv + (Discrete.equivalence Equiv.ulift : Discrete (ULift.{u} ι) ≌ _) _ /-- (Implementation Detail) Coproduct of schemes is isomorphic to the disjoint union. -/ noncomputable @@ -461,16 +463,16 @@ instance (R S : CommRingCatᵒᵖ) : IsIso (coprodComparison Scheme.Spec R S) := noncomputable instance : PreservesColimitsOfShape (Discrete WalkingPair) Scheme.Spec := ⟨fun {_} ↦ - have (X Y : CommRingCatᵒᵖ) := PreservesColimitPair.ofIsoCoprodComparison Scheme.Spec X Y - preservesColimitOfIsoDiagram _ (diagramIsoPair _).symm⟩ + have (X Y : CommRingCatᵒᵖ) := PreservesColimitPair.of_iso_coprod_comparison Scheme.Spec X Y + preservesColimit_of_iso_diagram _ (diagramIsoPair _).symm⟩ noncomputable instance : PreservesColimitsOfShape (Discrete PEmpty.{1}) Scheme.Spec := by have : IsEmpty (Scheme.Spec.obj (⊥_ CommRingCatᵒᵖ)) := @Function.isEmpty _ _ spec_punit_isEmpty (Scheme.Spec.mapIso (initialIsoIsInitial (initialOpOfTerminal CommRingCat.punitIsTerminal))).hom.base - have := preservesInitialOfIso Scheme.Spec (asIso (initial.to _)) - exact preservesColimitsOfShapePemptyOfPreservesInitial _ + have := preservesInitial_of_iso Scheme.Spec (asIso (initial.to _)) + exact preservesColimitsOfShape_pempty_of_preservesInitial _ noncomputable instance {J} [Fintype J] : PreservesColimitsOfShape (Discrete J) Scheme.Spec := @@ -479,7 +481,7 @@ instance {J} [Fintype J] : PreservesColimitsOfShape (Discrete J) Scheme.Spec := noncomputable instance {J : Type*} [Finite J] : PreservesColimitsOfShape (Discrete J) Scheme.Spec := letI := (nonempty_fintype J).some - preservesColimitsOfShapeOfEquiv (Discrete.equivalence (Fintype.equivFin _).symm) _ + preservesColimitsOfShape_of_equiv (Discrete.equivalence (Fintype.equivFin _).symm) _ /-- The canonical map `∐ Spec Rᵢ ⟶ Spec (Π Rᵢ)`. This is an isomorphism when the product is finite. -/ diff --git a/Mathlib/AlgebraicGeometry/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/OpenImmersion.lean index de7842550570e..834a151635fb5 100644 --- a/Mathlib/AlgebraicGeometry/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/OpenImmersion.lean @@ -454,11 +454,11 @@ instance forgetCreatesPullbackOfRight : CreatesLimit (cospan g f) forget := (PresheafedSpace.IsOpenImmersion.toScheme Y (pullback.fst g.toLRSHom f.toLRSHom).1) (eqToIso (by simp) ≪≫ HasLimit.isoOfNatIso (diagramIsoCospan _).symm) -instance forgetPreservesOfLeft : PreservesLimit (cospan f g) forget := - CategoryTheory.preservesLimitOfCreatesLimitAndHasLimit _ _ +instance forget_preservesOfLeft : PreservesLimit (cospan f g) forget := + CategoryTheory.preservesLimit_of_createsLimit_and_hasLimit _ _ -instance forgetPreservesOfRight : PreservesLimit (cospan g f) forget := - preservesPullbackSymmetry _ _ _ +instance forget_preservesOfRight : PreservesLimit (cospan g f) forget := + preservesPullback_symmetry _ _ _ instance hasPullback_of_left : HasPullback f g := hasLimit_of_created (cospan f g) forget @@ -487,17 +487,17 @@ instance pullback_to_base [IsOpenImmersion g] : -- provided but still class inference fail to find this exact LocallyRingedSpace.IsOpenImmersion.comp (H := inferInstance) _ _ -instance forgetToTopPreservesOfLeft : PreservesLimit (cospan f g) Scheme.forgetToTop := by +instance forgetToTop_preserves_of_left : PreservesLimit (cospan f g) Scheme.forgetToTop := by delta Scheme.forgetToTop - refine @Limits.compPreservesLimit _ _ _ _ _ _ (K := cospan f g) _ _ (F := forget) + refine @Limits.comp_preservesLimit _ _ _ _ _ _ (K := cospan f g) _ _ (F := forget) (G := LocallyRingedSpace.forgetToTop) ?_ ?_ · infer_instance - refine @preservesLimitOfIsoDiagram _ _ _ _ _ _ _ _ _ (diagramIsoCospan.{u} _).symm ?_ + refine @preservesLimit_of_iso_diagram _ _ _ _ _ _ _ _ _ (diagramIsoCospan.{u} _).symm ?_ dsimp [LocallyRingedSpace.forgetToTop] infer_instance -instance forgetToTopPreservesOfRight : PreservesLimit (cospan g f) Scheme.forgetToTop := - preservesPullbackSymmetry _ _ _ +instance forgetToTop_preserves_of_right : PreservesLimit (cospan g f) Scheme.forgetToTop := + preservesPullback_symmetry _ _ _ theorem range_pullback_snd_of_left : Set.range (pullback.snd f g).base = (g ⁻¹ᵁ f.opensRange).1 := by diff --git a/Mathlib/AlgebraicTopology/FundamentalGroupoid/Product.lean b/Mathlib/AlgebraicTopology/FundamentalGroupoid/Product.lean index e6ebf8e4a8192..2877b1ad5f01d 100644 --- a/Mathlib/AlgebraicTopology/FundamentalGroupoid/Product.lean +++ b/Mathlib/AlgebraicTopology/FundamentalGroupoid/Product.lean @@ -107,9 +107,9 @@ instance : IsIso (piTopToPiCone X) := Limits.Cones.cone_iso_of_hom_iso (piTopToPiCone X) /-- The fundamental groupoid functor preserves products -/ -def preservesProduct : Limits.PreservesLimit (Discrete.functor X) π := by +lemma preservesProduct : Limits.PreservesLimit (Discrete.functor X) π := by -- Porting note: check universe parameters here - apply Limits.preservesLimitOfPreservesLimitCone (TopCat.piFanIsLimit.{u,u} X) + apply Limits.preservesLimit_of_preserves_limit_cone (TopCat.piFanIsLimit.{u,u} X) apply (Limits.IsLimit.ofConeEquiv (coneDiscreteComp X)).toFun simp only [coneDiscreteComp_obj_mapCone] apply Limits.IsLimit.ofIsoLimit _ (asIso (piTopToPiCone X)).symm diff --git a/Mathlib/CategoryTheory/Abelian/Exact.lean b/Mathlib/CategoryTheory/Abelian/Exact.lean index 04967ba6b7f80..badc3ad7a0d10 100644 --- a/Mathlib/CategoryTheory/Abelian/Exact.lean +++ b/Mathlib/CategoryTheory/Abelian/Exact.lean @@ -244,10 +244,10 @@ theorem preservesEpimorphisms_of_map_exact : L.PreservesEpimorphisms where exact ShortComplex.isoMk (Iso.refl _) (Iso.refl _) (Iso.refl _) /-- A functor which preserves the exactness of short complexes preserves homology. -/ -def preservesHomologyOfMapExact : L.PreservesHomology where +lemma preservesHomology_of_map_exact : L.PreservesHomology where preservesCokernels X Y f := by have := preservesEpimorphisms_of_map_exact _ hL - apply preservesColimitOfPreservesColimitCocone (cokernelIsCokernel f) + apply preservesColimit_of_preserves_colimit_cocone (cokernelIsCokernel f) apply (CokernelCofork.isColimitMapCoconeEquiv _ L).2 have : Epi ((ShortComplex.mk _ _ (cokernel.condition f)).map L).g := by dsimp @@ -256,7 +256,7 @@ def preservesHomologyOfMapExact : L.PreservesHomology where (ShortComplex.exact_of_g_is_cokernel _ (cokernelIsCokernel f))).gIsCokernel preservesKernels X Y f := by have := preservesMonomorphisms_of_map_exact _ hL - apply preservesLimitOfPreservesLimitCone (kernelIsKernel f) + apply preservesLimit_of_preserves_limit_cone (kernelIsKernel f) apply (KernelFork.isLimitMapConeEquiv _ L).2 have : Mono ((ShortComplex.mk _ _ (kernel.condition f)).map L).f := by dsimp @@ -274,10 +274,10 @@ end section /-- A functor preserving zero morphisms, monos, and cokernels preserves homology. -/ -def preservesHomologyOfPreservesMonosAndCokernels [PreservesZeroMorphisms L] +lemma preservesHomology_of_preservesMonos_and_cokernels [PreservesZeroMorphisms L] [PreservesMonomorphisms L] [∀ {X Y} (f : X ⟶ Y), PreservesColimit (parallelPair f 0) L] : PreservesHomology L := by - apply preservesHomologyOfMapExact + apply preservesHomology_of_map_exact intro S hS let φ : (ShortComplex.mk _ _ (Abelian.comp_coimage_π_eq_zero S.zero)).map L ⟶ S.map L := { τ₁ := 𝟙 _ @@ -291,10 +291,10 @@ def preservesHomologyOfPreservesMonosAndCokernels [PreservesZeroMorphisms L] exact CokernelCofork.mapIsColimit _ ((S.exact_iff_exact_coimage_π).1 hS).gIsCokernel L /-- A functor preserving zero morphisms, epis, and kernels preserves homology. -/ -def preservesHomologyOfPreservesEpisAndKernels [PreservesZeroMorphisms L] +lemma preservesHomology_of_preservesEpis_and_kernels [PreservesZeroMorphisms L] [PreservesEpimorphisms L] [∀ {X Y} (f : X ⟶ Y), PreservesLimit (parallelPair f 0) L] : PreservesHomology L := by - apply preservesHomologyOfMapExact + apply preservesHomology_of_map_exact intro S hS let φ : S.map L ⟶ (ShortComplex.mk _ _ (Abelian.image_ι_comp_eq_zero S.zero)).map L := { τ₁ := L.map (Abelian.factorThruImage S.f) diff --git a/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms.lean b/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms.lean index 659a9efa8a888..e11ac111c917c 100644 --- a/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms.lean +++ b/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms.lean @@ -107,24 +107,24 @@ open CoproductsFromFiniteFiltered variable {α : Type w} variable [HasZeroMorphisms C] [HasFiniteBiproducts C] [HasFiniteLimits C] -instance preservesFiniteLimitsLiftToFinset : PreservesFiniteLimits (liftToFinset C α) := - preservesFiniteLimitsOfEvaluation _ fun I => +instance preservesFiniteLimits_liftToFinset : PreservesFiniteLimits (liftToFinset C α) := + preservesFiniteLimits_of_evaluation _ fun I => letI : PreservesFiniteLimits (colim (J := Discrete I) (C := C)) := - preservesFiniteLimitsOfNatIso HasBiproductsOfShape.colimIsoLim.symm + preservesFiniteLimits_of_natIso HasBiproductsOfShape.colimIsoLim.symm letI : PreservesFiniteLimits ((whiskeringLeft (Discrete I) (Discrete α) C).obj (Discrete.functor fun x ↦ ↑x)) := - ⟨fun J _ _ => whiskeringLeftPreservesLimitsOfShape J _⟩ + ⟨fun J _ _ => whiskeringLeft_preservesLimitsOfShape J _⟩ letI : PreservesFiniteLimits ((whiskeringLeft (Discrete I) (Discrete α) C).obj (Discrete.functor (·.val)) ⋙ colim) := - compPreservesFiniteLimits _ _ - preservesFiniteLimitsOfNatIso (liftToFinsetEvaluationIso I).symm + comp_preservesFiniteLimits _ _ + preservesFiniteLimits_of_natIso (liftToFinsetEvaluationIso I).symm /-- A category with finite biproducts and finite limits is AB4 if it is AB5. -/ -def AB4.ofAB5 [HasFilteredColimits C] [AB5 C] : AB4 C where +lemma AB4.of_AB5 [HasFilteredColimits C] [AB5 C] : AB4 C where preservesFiniteLimits J := letI : PreservesFiniteLimits (liftToFinset C J ⋙ colim) := - compPreservesFiniteLimits _ _ - preservesFiniteLimitsOfNatIso (liftToFinsetColimIso) + comp_preservesFiniteLimits _ _ + preservesFiniteLimits_of_natIso (liftToFinsetColimIso) end diff --git a/Mathlib/CategoryTheory/Abelian/Injective.lean b/Mathlib/CategoryTheory/Abelian/Injective.lean index da13360d678e4..b03c8db8b549c 100644 --- a/Mathlib/CategoryTheory/Abelian/Injective.lean +++ b/Mathlib/CategoryTheory/Abelian/Injective.lean @@ -28,15 +28,15 @@ namespace CategoryTheory variable {C : Type u} [Category.{v} C] [Abelian C] /-- The preadditive Yoneda functor on `J` preserves homology if `J` is injective. -/ -instance preservesHomologyPreadditiveYonedaObjOfInjective (J : C) [hJ : Injective J] : +instance preservesHomology_preadditiveYonedaObj_of_injective (J : C) [hJ : Injective J] : (preadditiveYonedaObj J).PreservesHomology := by letI := (injective_iff_preservesEpimorphisms_preadditive_yoneda_obj' J).mp hJ - apply Functor.preservesHomologyOfPreservesEpisAndKernels + apply Functor.preservesHomology_of_preservesEpis_and_kernels /-- The preadditive Yoneda functor on `J` preserves colimits if `J` is injective. -/ -instance preservesFiniteColimitsPreadditiveYonedaObjOfInjective (J : C) [hP : Injective J] : +instance preservesFiniteColimits_preadditiveYonedaObj_of_injective (J : C) [hP : Injective J] : PreservesFiniteColimits (preadditiveYonedaObj J) := by - apply Functor.preservesFiniteColimitsOfPreservesHomology + apply Functor.preservesFiniteColimits_of_preservesHomology /-- An object is injective if its preadditive Yoneda functor preserves finite colimits. -/ theorem injective_of_preservesFiniteColimits_preadditiveYonedaObj (J : C) diff --git a/Mathlib/CategoryTheory/Abelian/Projective.lean b/Mathlib/CategoryTheory/Abelian/Projective.lean index 1eb7a8843324e..af417648361c8 100644 --- a/Mathlib/CategoryTheory/Abelian/Projective.lean +++ b/Mathlib/CategoryTheory/Abelian/Projective.lean @@ -26,18 +26,18 @@ open Limits Projective Opposite variable {C : Type u} [Category.{v} C] [Abelian C] /-- The preadditive Co-Yoneda functor on `P` preserves homology if `P` is projective. -/ -noncomputable instance preservesHomologyPreadditiveCoyonedaObjOfProjective +noncomputable instance preservesHomology_preadditiveCoyonedaObj_of_projective (P : C) [hP : Projective P] : (preadditiveCoyonedaObj (op P)).PreservesHomology := by haveI := (projective_iff_preservesEpimorphisms_preadditiveCoyoneda_obj' P).mp hP haveI := @Functor.preservesEpimorphisms_of_preserves_of_reflects _ _ _ _ _ _ _ _ this _ - apply Functor.preservesHomologyOfPreservesEpisAndKernels + apply Functor.preservesHomology_of_preservesEpis_and_kernels /-- The preadditive Co-Yoneda functor on `P` preserves finite colimits if `P` is projective. -/ -noncomputable instance preservesFiniteColimitsPreadditiveCoyonedaObjOfProjective +noncomputable instance preservesFiniteColimits_preadditiveCoyonedaObj_of_projective (P : C) [hP : Projective P] : PreservesFiniteColimits (preadditiveCoyonedaObj (op P)) := by - apply Functor.preservesFiniteColimitsOfPreservesHomology + apply Functor.preservesFiniteColimits_of_preservesHomology /-- An object is projective if its preadditive Co-Yoneda functor preserves finite colimits. -/ theorem projective_of_preservesFiniteColimits_preadditiveCoyonedaObj (P : C) diff --git a/Mathlib/CategoryTheory/Abelian/Transfer.lean b/Mathlib/CategoryTheory/Abelian/Transfer.lean index 5a3c93413c248..edda974b2f59c 100644 --- a/Mathlib/CategoryTheory/Abelian/Transfer.lean +++ b/Mathlib/CategoryTheory/Abelian/Transfer.lean @@ -59,7 +59,7 @@ theorem hasKernels [PreservesFiniteLimits G] (i : F ⋙ G ≅ 𝟭 C) : HasKerne /-- No point making this an instance, as it requires `i` and `adj`. -/ theorem hasCokernels (i : F ⋙ G ≅ 𝟭 C) (adj : G ⊣ F) : HasCokernels C := { has_colimit := fun f => by - have : PreservesColimits G := adj.leftAdjointPreservesColimits + have : PreservesColimits G := adj.leftAdjoint_preservesColimits have := NatIso.naturality_1 i f simp? at this says simp only [Functor.id_obj, Functor.comp_obj, Functor.comp_map, Functor.id_map] at this @@ -74,7 +74,7 @@ def cokernelIso (i : F ⋙ G ≅ 𝟭 C) (adj : G ⊣ F) {X Y : C} (f : X ⟶ Y) G.obj (cokernel (F.map f)) ≅ cokernel f := by -- We have to write an explicit `PreservesColimits` type here, -- as `leftAdjointPreservesColimits` has universe variables. - have : PreservesColimits G := adj.leftAdjointPreservesColimits + have : PreservesColimits G := adj.leftAdjoint_preservesColimits calc G.obj (cokernel (F.map f)) ≅ cokernel (G.map (F.map f)) := (asIso (cokernelComparison _ G)).symm @@ -87,7 +87,7 @@ variable [Limits.HasKernels C] [PreservesFiniteLimits G] /-- Auxiliary construction for `coimageIsoImage` -/ def coimageIsoImageAux (i : F ⋙ G ≅ 𝟭 C) (adj : G ⊣ F) {X Y : C} (f : X ⟶ Y) : kernel (G.map (cokernel.π (F.map f))) ≅ kernel (cokernel.π f) := by - have : PreservesColimits G := adj.leftAdjointPreservesColimits + have : PreservesColimits G := adj.leftAdjoint_preservesColimits calc kernel (G.map (cokernel.π (F.map f))) ≅ kernel (cokernel.π (G.map (F.map f)) ≫ cokernelComparison (F.map f) G) := @@ -113,7 +113,7 @@ We still need to check that this agrees with the canonical morphism. -/ def coimageIsoImage (i : F ⋙ G ≅ 𝟭 C) (adj : G ⊣ F) {X Y : C} (f : X ⟶ Y) : Abelian.coimage f ≅ Abelian.image f := by - have : PreservesLimits F := adj.rightAdjointPreservesLimits + have : PreservesLimits F := adj.rightAdjoint_preservesLimits calc Abelian.coimage f ≅ cokernel (kernel.ι f) := Iso.refl _ _ ≅ G.obj (cokernel (F.map (kernel.ι f))) := (cokernelIso _ _ i adj _).symm diff --git a/Mathlib/CategoryTheory/Adhesive.lean b/Mathlib/CategoryTheory/Adhesive.lean index 51235c1937b60..1337051cd7b2a 100644 --- a/Mathlib/CategoryTheory/Adhesive.lean +++ b/Mathlib/CategoryTheory/Adhesive.lean @@ -293,9 +293,9 @@ theorem adhesive_of_preserves_and_reflects_isomorphism (F : C ⥤ D) [F.ReflectsIsomorphisms] : Adhesive C := by haveI : ReflectsLimitsOfShape WalkingCospan F := - reflectsLimitsOfShapeOfReflectsIsomorphisms + reflectsLimitsOfShape_of_reflectsIsomorphisms haveI : ReflectsColimitsOfShape WalkingSpan F := - reflectsColimitsOfShapeOfReflectsIsomorphisms + reflectsColimitsOfShape_of_reflectsIsomorphisms exact adhesive_of_preserves_and_reflects F theorem adhesive_of_reflective [HasPullbacks D] [Adhesive C] [HasPullbacks C] [HasPushouts C] @@ -303,8 +303,8 @@ theorem adhesive_of_reflective [HasPullbacks D] [Adhesive C] [HasPullbacks C] [H {Gl : C ⥤ D} {Gr : D ⥤ C} (adj : Gl ⊣ Gr) [Gr.Full] [Gr.Faithful] [PreservesLimitsOfShape WalkingCospan Gl] : Adhesive D := by - have := adj.leftAdjointPreservesColimits - have := adj.rightAdjointPreservesLimits + have := adj.leftAdjoint_preservesColimits + have := adj.rightAdjoint_preservesLimits apply Adhesive.mk (hasPushout_of_mono_left := H₂) intro W X Y Z f g h i _ H have := Adhesive.van_kampen (IsPushout.of_hasPushout (Gr.map f) (Gr.map g)) diff --git a/Mathlib/CategoryTheory/Adjunction/Limits.lean b/Mathlib/CategoryTheory/Adjunction/Limits.lean index f57ec0e073cee..dd86591976534 100644 --- a/Mathlib/CategoryTheory/Adjunction/Limits.lean +++ b/Mathlib/CategoryTheory/Adjunction/Limits.lean @@ -9,14 +9,14 @@ import Mathlib.CategoryTheory.Limits.Creates /-! # Adjunctions and limits -A left adjoint preserves colimits (`CategoryTheory.Adjunction.leftAdjointPreservesColimits`), -and a right adjoint preserves limits (`CategoryTheory.Adjunction.rightAdjointPreservesLimits`). +A left adjoint preserves colimits (`CategoryTheory.Adjunction.leftAdjoint_preservesColimits`), +and a right adjoint preserves limits (`CategoryTheory.Adjunction.rightAdjoint_preservesLimits`). Equivalences create and reflect (co)limits. -(`CategoryTheory.Adjunction.isEquivalenceCreatesLimits`, -`CategoryTheory.Adjunction.isEquivalenceCreatesColimits`, -`CategoryTheory.Adjunction.isEquivalenceReflectsLimits`, -`CategoryTheory.Adjunction.isEquivalenceReflectsColimits`,) +(`CategoryTheory.Functor.createsLimitsOfIsEquivalence`, +`CategoryTheory.Functor.createsColimitsOfIsEquivalence`, +`CategoryTheory.Functor.reflectsLimits_of_isEquivalence`, +`CategoryTheory.Functor.reflectsColimits_of_isEquivalence`.) In `CategoryTheory.Adjunction.coconesIso` we show that when `F ⊣ G`, @@ -78,40 +78,53 @@ def functorialityAdjunction : Cocones.functoriality K F ⊣ functorialityRightAd unit := functorialityUnit adj K counit := functorialityCounit adj K +include adj in /-- A left adjoint preserves colimits. See . -/ -def leftAdjointPreservesColimits : PreservesColimitsOfSize.{v, u} F where +lemma leftAdjoint_preservesColimits : PreservesColimitsOfSize.{v, u} F where preservesColimitsOfShape := { preservesColimit := { preserves := fun hc => - IsColimit.isoUniqueCoconeMorphism.inv fun _ => + ⟨IsColimit.isoUniqueCoconeMorphism.inv fun _ => @Equiv.unique _ _ (IsColimit.isoUniqueCoconeMorphism.hom hc _) - ((adj.functorialityAdjunction _).homEquiv _ _) } } + ((adj.functorialityAdjunction _).homEquiv _ _)⟩ } } + +include adj in +@[deprecated (since := "2024-11-19")] +lemma leftAdjointPreservesColimits : PreservesColimitsOfSize.{v, u} F := + adj.leftAdjoint_preservesColimits noncomputable -instance colimPreservesColimits [HasColimitsOfShape J C] : +instance colim_preservesColimits [HasColimitsOfShape J C] : PreservesColimits (colim (J := J) (C := C)) := - colimConstAdj.leftAdjointPreservesColimits + colimConstAdj.leftAdjoint_preservesColimits -- see Note [lower instance priority] -noncomputable instance (priority := 100) isEquivalencePreservesColimits +noncomputable instance (priority := 100) isEquivalence_preservesColimits (E : C ⥤ D) [E.IsEquivalence] : PreservesColimitsOfSize.{v, u} E := - leftAdjointPreservesColimits E.adjunction + leftAdjoint_preservesColimits E.adjunction -- see Note [lower instance priority] -noncomputable instance (priority := 100) isEquivalenceReflectsColimits +noncomputable instance (priority := 100) + _root_.CategoryTheory.Functor.reflectsColimits_of_isEquivalence (E : D ⥤ C) [E.IsEquivalence] : ReflectsColimitsOfSize.{v, u} E where reflectsColimitsOfShape := { reflectsColimit := { reflects := fun t => - (isColimitOfPreserves E.inv t).mapCoconeEquiv E.asEquivalence.unitIso.symm } } + ⟨(isColimitOfPreserves E.inv t).mapCoconeEquiv E.asEquivalence.unitIso.symm⟩ } } + +@[deprecated (since := "2024-11-18")] +lemma isEquivalenceReflectsColimits (E : D ⥤ C) [E.IsEquivalence] : + ReflectsColimitsOfSize.{v, u} E := + Functor.reflectsColimits_of_isEquivalence E -- see Note [lower instance priority] -noncomputable instance (priority := 100) isEquivalenceCreatesColimits (H : D ⥤ C) +noncomputable instance (priority := 100) + _root_.CategoryTheory.Functor.createsColimitsOfIsEquivalence (H : D ⥤ C) [H.IsEquivalence] : CreatesColimitsOfSize.{v, u} H where CreatesColimitsOfShape := @@ -120,16 +133,19 @@ noncomputable instance (priority := 100) isEquivalenceCreatesColimits (H : D ⥤ { liftedCocone := mapCoconeInv H c validLift := mapCoconeMapCoconeInv H c } } } +@[deprecated (since := "2024-11-18")] alias isEquivalenceCreatesColimits := + Functor.createsColimitsOfIsEquivalence + -- verify the preserve_colimits instance works as expected: noncomputable example (E : C ⥤ D) [E.IsEquivalence] (c : Cocone K) (h : IsColimit c) : IsColimit (E.mapCocone c) := - PreservesColimit.preserves h + isColimitOfPreserves E h theorem hasColimit_comp_equivalence (E : C ⥤ D) [E.IsEquivalence] [HasColimit K] : HasColimit (K ⋙ E) := HasColimit.mk { cocone := E.mapCocone (colimit.cocone K) - isColimit := PreservesColimit.preserves (colimit.isColimit K) } + isColimit := isColimitOfPreserves _ (colimit.isColimit K) } theorem hasColimit_of_comp_equivalence (E : C ⥤ D) [E.IsEquivalence] [HasColimit (K ⋙ E)] : HasColimit K := @@ -186,40 +202,52 @@ def functorialityAdjunction' : functorialityLeftAdjoint adj K ⊣ Cones.functori unit := functorialityUnit' adj K counit := functorialityCounit' adj K +include adj in /-- A right adjoint preserves limits. See . -/ -def rightAdjointPreservesLimits : PreservesLimitsOfSize.{v, u} G where +lemma rightAdjoint_preservesLimits : PreservesLimitsOfSize.{v, u} G where preservesLimitsOfShape := { preservesLimit := { preserves := fun hc => - IsLimit.isoUniqueConeMorphism.inv fun _ => + ⟨IsLimit.isoUniqueConeMorphism.inv fun _ => @Equiv.unique _ _ (IsLimit.isoUniqueConeMorphism.hom hc _) - ((adj.functorialityAdjunction' _).homEquiv _ _).symm } } + ((adj.functorialityAdjunction' _).homEquiv _ _).symm⟩ } } -noncomputable -instance limPreservesLimits [HasLimitsOfShape J C] : +include adj in +@[deprecated (since := "2024-11-19")] +lemma rightAdjointPreservesLimits : PreservesLimitsOfSize.{v, u} G := + adj.rightAdjoint_preservesLimits + +instance lim_preservesLimits [HasLimitsOfShape J C] : PreservesLimits (lim (J := J) (C := C)) := - constLimAdj.rightAdjointPreservesLimits + constLimAdj.rightAdjoint_preservesLimits -- see Note [lower instance priority] -noncomputable instance (priority := 100) isEquivalencePreservesLimits +instance (priority := 100) isEquivalencePreservesLimits (E : D ⥤ C) [E.IsEquivalence] : PreservesLimitsOfSize.{v, u} E := - rightAdjointPreservesLimits E.asEquivalence.symm.toAdjunction + rightAdjoint_preservesLimits E.asEquivalence.symm.toAdjunction -- see Note [lower instance priority] -noncomputable instance (priority := 100) isEquivalenceReflectsLimits +noncomputable instance (priority := 100) + _root_.CategoryTheory.Functor.reflectsLimits_of_isEquivalence (E : D ⥤ C) [E.IsEquivalence] : ReflectsLimitsOfSize.{v, u} E where reflectsLimitsOfShape := { reflectsLimit := { reflects := fun t => - (isLimitOfPreserves E.inv t).mapConeEquiv E.asEquivalence.unitIso.symm } } + ⟨(isLimitOfPreserves E.inv t).mapConeEquiv E.asEquivalence.unitIso.symm⟩ } } + +@[deprecated (since := "2024-11-18")] +lemma isEquivalenceReflectsLimits (E : D ⥤ C) [E.IsEquivalence] : + ReflectsLimitsOfSize.{v, u} E := + Functor.reflectsLimits_of_isEquivalence E -- see Note [lower instance priority] -noncomputable instance (priority := 100) isEquivalenceCreatesLimits (H : D ⥤ C) [H.IsEquivalence] : +noncomputable instance (priority := 100) + _root_.CategoryTheory.Functor.createsLimitsOfIsEquivalence (H : D ⥤ C) [H.IsEquivalence] : CreatesLimitsOfSize.{v, u} H where CreatesLimitsOfShape := { CreatesLimit := @@ -227,15 +255,18 @@ noncomputable instance (priority := 100) isEquivalenceCreatesLimits (H : D ⥤ C { liftedCone := mapConeInv H c validLift := mapConeMapConeInv H c } } } +@[deprecated (since := "2024-11-18")] alias isEquivalenceCreatesLimits := + Functor.createsLimitsOfIsEquivalence + -- verify the preserve_limits instance works as expected: noncomputable example (E : D ⥤ C) [E.IsEquivalence] (c : Cone K) (h : IsLimit c) : IsLimit (E.mapCone c) := - PreservesLimit.preserves h + isLimitOfPreserves E h theorem hasLimit_comp_equivalence (E : D ⥤ C) [E.IsEquivalence] [HasLimit K] : HasLimit (K ⋙ E) := HasLimit.mk { cone := E.mapCone (limit.cone K) - isLimit := PreservesLimit.preserves (limit.isLimit K) } + isLimit := isLimitOfPreserves _ (limit.isLimit K) } theorem hasLimit_of_comp_equivalence (E : D ⥤ C) [E.IsEquivalence] [HasLimit (K ⋙ E)] : HasLimit K := @@ -330,12 +361,12 @@ variable {J C D : Type*} [Category J] [Category C] [Category D] (F : C ⥤ D) noncomputable instance [IsLeftAdjoint F] : PreservesColimitsOfShape J F := - (Adjunction.ofIsLeftAdjoint F).leftAdjointPreservesColimits.preservesColimitsOfShape + (Adjunction.ofIsLeftAdjoint F).leftAdjoint_preservesColimits.preservesColimitsOfShape noncomputable instance [IsLeftAdjoint F] : PreservesColimitsOfSize.{v, u} F where noncomputable instance [IsRightAdjoint F] : PreservesLimitsOfShape J F := - (Adjunction.ofIsRightAdjoint F).rightAdjointPreservesLimits.preservesLimitsOfShape + (Adjunction.ofIsRightAdjoint F).rightAdjoint_preservesLimits.preservesLimitsOfShape noncomputable instance [IsRightAdjoint F] : PreservesLimitsOfSize.{v, u} F where diff --git a/Mathlib/CategoryTheory/Category/Cat/Limit.lean b/Mathlib/CategoryTheory/Category/Cat/Limit.lean index 92300512f30e6..e16a6957e1e3f 100644 --- a/Mathlib/CategoryTheory/Category/Cat/Limit.lean +++ b/Mathlib/CategoryTheory/Category/Cat/Limit.lean @@ -146,7 +146,7 @@ instance : HasLimits Cat.{v, v} where instance : PreservesLimits Cat.objects.{v, v} where preservesLimitsOfShape := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone (HasLimits.limitConeIsLimit F) + preservesLimit_of_preserves_limit_cone (HasLimits.limitConeIsLimit F) (Limits.IsLimit.ofIsoLimit (limit.isLimit (F ⋙ Cat.objects)) (Cones.ext (by rfl) (by aesop_cat))) } diff --git a/Mathlib/CategoryTheory/ChosenFiniteProducts.lean b/Mathlib/CategoryTheory/ChosenFiniteProducts.lean index 80e1a62db2144..0172d99e2112c 100644 --- a/Mathlib/CategoryTheory/ChosenFiniteProducts.lean +++ b/Mathlib/CategoryTheory/ChosenFiniteProducts.lean @@ -246,15 +246,16 @@ lemma map_toUnit_comp_terminalCompariso (A : C) : open Limits /-- If `terminalComparison F` is an Iso, then `F` preserves terminal objects. -/ -noncomputable def preservesLimitEmptyOfIsIsoTerminalComparison [IsIso (terminalComparison F)] : +lemma preservesLimit_empty_of_isIso_terminalComparison [IsIso (terminalComparison F)] : PreservesLimit (Functor.empty.{0} C) F := by - apply preservesLimitOfPreservesLimitCone terminal.isLimit + apply preservesLimit_of_preserves_limit_cone terminal.isLimit apply isLimitChangeEmptyCone D terminal.isLimit exact asIso (terminalComparison F)|>.symm /-- If `F` preserves terminal objects, then `terminalComparison F` is an isomorphism. -/ -def preservesTerminalIso [h : PreservesLimit (Functor.empty.{0} C) F] : F.obj (𝟙_ C) ≅ 𝟙_ D := - (isLimitChangeEmptyCone D (h.preserves terminal.isLimit) (asEmptyCone (F.obj (𝟙_ C))) +noncomputable def preservesTerminalIso [h : PreservesLimit (Functor.empty.{0} C) F] : + F.obj (𝟙_ C) ≅ 𝟙_ D := + (isLimitChangeEmptyCone D (isLimitOfPreserves _ terminal.isLimit) (asEmptyCone (F.obj (𝟙_ C))) (Iso.refl _)).conePointUniqueUpToIso terminal.isLimit @[simp] @@ -415,14 +416,14 @@ variable [PreservesLimit (pair A B) F] /-- If `F` preserves the limit of the pair `(A, B)`, then the binary fan given by `(F.map fst A B, F.map (snd A B))` is a limit cone. -/ -def isLimitChosenFiniteProductsOfPreservesLimits : +noncomputable def isLimitChosenFiniteProductsOfPreservesLimits : IsLimit <| BinaryFan.mk (F.map (fst A B)) (F.map (snd A B)) := mapIsLimitOfPreservesOfIsLimit F (fst _ _) (snd _ _) <| (product A B).isLimit.ofIsoLimit <| isoBinaryFanMk (product A B).cone /-- If `F` preserves the limit of the pair `(A, B)`, then `prodComparison F A B` is an isomorphism. -/ -def prodComparisonIso : F.obj (A ⊗ B) ≅ F.obj A ⊗ F.obj B := +noncomputable def prodComparisonIso : F.obj (A ⊗ B) ≅ F.obj A ⊗ F.obj B := IsLimit.conePointUniqueUpToIso (isLimitChosenFiniteProductsOfPreservesLimits F A B) (product _ _).isLimit @@ -456,9 +457,10 @@ end PreservesLimitPairs section ProdComparisonIso /-- If `prodComparison F A B` is an isomorphism, then `F` preserves the limit of `pair A B`. -/ -noncomputable def preservesLimitPairOfIsIsoProdComparison (A B : C) [IsIso (prodComparison F A B)] : +lemma preservesLimit_pair_of_isIso_prodComparison (A B : C) + [IsIso (prodComparison F A B)] : PreservesLimit (pair A B) F := by - apply preservesLimitOfPreservesLimitCone (product A B).isLimit + apply preservesLimit_of_preserves_limit_cone (product A B).isLimit refine IsLimit.equivOfNatIsoOfIso (pairComp A B F) _ ((product (F.obj A) (F.obj B)).cone.extend (prodComparison F A B)) (BinaryFan.ext (by exact Iso.refl _) ?_ ?_) |>.invFun @@ -470,12 +472,12 @@ noncomputable def preservesLimitPairOfIsIsoProdComparison (A B : C) [IsIso (prod /-- If `prodComparison F A B` is an isomorphism for all `A B` then `F` preserves limits of shape `Discrete (WalkingPair)`. -/ -noncomputable def preservesLimitsOfShapeDiscreteWalkingPairOfIsoProdComparison +lemma preservesLimitsOfShape_discrete_walkingPair_of_isIso_prodComparison [∀ A B, IsIso (prodComparison F A B)] : PreservesLimitsOfShape (Discrete WalkingPair) F := by constructor intro K - refine @preservesLimitOfIsoDiagram _ _ _ _ _ _ _ _ _ (diagramIsoPair K).symm ?_ - apply preservesLimitPairOfIsIsoProdComparison + refine @preservesLimit_of_iso_diagram _ _ _ _ _ _ _ _ _ (diagramIsoPair K).symm ?_ + apply preservesLimit_pair_of_isIso_prodComparison end ProdComparisonIso diff --git a/Mathlib/CategoryTheory/ChosenFiniteProducts/FunctorCategory.lean b/Mathlib/CategoryTheory/ChosenFiniteProducts/FunctorCategory.lean index 3a4eaa80fc91f..68e9344407ae8 100644 --- a/Mathlib/CategoryTheory/ChosenFiniteProducts/FunctorCategory.lean +++ b/Mathlib/CategoryTheory/ChosenFiniteProducts/FunctorCategory.lean @@ -154,11 +154,11 @@ lemma associator_inv_app (F₁ F₂ F₃ : J ⥤ C) (j : J) : noncomputable instance {K : Type*} [Category K] [HasColimitsOfShape K C] [∀ X : C, PreservesColimitsOfShape K (tensorLeft X)] {F : J ⥤ C} : PreservesColimitsOfShape K (tensorLeft F) := by - apply preservesColimitsOfShapeOfEvaluation + apply preservesColimitsOfShape_of_evaluation intro k haveI : tensorLeft F ⋙ (evaluation J C).obj k ≅ (evaluation J C).obj k ⋙ tensorLeft (F.obj k) := NatIso.ofComponents (fun _ ↦ Iso.refl _) - exact preservesColimitsOfShapeOfNatIso this.symm + exact preservesColimitsOfShape_of_natIso this.symm end Monoidal diff --git a/Mathlib/CategoryTheory/Closed/Cartesian.lean b/Mathlib/CategoryTheory/Closed/Cartesian.lean index 90b541c12ab47..6b72829dd536d 100644 --- a/Mathlib/CategoryTheory/Closed/Cartesian.lean +++ b/Mathlib/CategoryTheory/Closed/Cartesian.lean @@ -148,7 +148,7 @@ theorem coev_ev : (coev A).app (A ⟹ B) ≫ (exp A).map ((ev A).app B) = 𝟙 ( end exp instance : PreservesColimits (tensorLeft A) := - (ihom.adjunction A).leftAdjointPreservesColimits + (ihom.adjunction A).leftAdjoint_preservesColimits variable {A} diff --git a/Mathlib/CategoryTheory/Closed/Ideal.lean b/Mathlib/CategoryTheory/Closed/Ideal.lean index 729c2ce13fd84..7421ed490bb18 100644 --- a/Mathlib/CategoryTheory/Closed/Ideal.lean +++ b/Mathlib/CategoryTheory/Closed/Ideal.lean @@ -120,7 +120,7 @@ def reflectiveChosenFiniteProducts [ChosenFiniteProducts C] [Reflective i] : ((reflector i).map (fst (i.obj X) (i.obj Y)) ≫ (reflectorAdjunction i).counit.app _) ((reflector i).map (snd (i.obj X) (i.obj Y)) ≫ (reflectorAdjunction i).counit.app _) isLimit := by - apply (by infer_instance : ReflectsLimit _ i).reflects + apply isLimitOfReflects i apply IsLimit.equivOfNatIsoOfIso (pairComp X Y _) _ _ _|>.invFun (product (i.obj X) (i.obj Y)).isLimit fapply BinaryFan.ext @@ -141,7 +141,7 @@ def reflectiveChosenFiniteProducts [ChosenFiniteProducts C] [Reflective i] : terminal := { cone := Limits.asEmptyCone <| (reflector i).obj (𝟙_ C) isLimit := by - apply (by infer_instance : ReflectsLimit _ i).reflects + apply isLimitOfReflects i apply isLimitChangeEmptyCone _ ChosenFiniteProducts.terminal.isLimit letI : IsIso ((reflectorAdjunction i).unit.app (𝟙_ C)) := by apply Functor.essImage.unit_isIso @@ -195,15 +195,15 @@ def cartesianClosedOfReflective : CartesianClosed D where · symm refine NatIso.ofComponents (fun X => ?_) (fun f => ?_) · haveI := - Adjunction.rightAdjointPreservesLimits.{0, 0} (reflectorAdjunction i) + Adjunction.rightAdjoint_preservesLimits.{0, 0} (reflectorAdjunction i) apply asIso (prodComparison i B X) · dsimp [asIso] rw [prodComparison_natural_whiskerLeft] · apply (exponentialIdealReflective i _).symm } -- It's annoying that I need to do this. -attribute [-instance] CategoryTheory.preservesLimitOfCreatesLimitAndHasLimit - CategoryTheory.preservesLimitOfShapeOfCreatesLimitsOfShapeAndHasLimitsOfShape +attribute [-instance] CategoryTheory.preservesLimit_of_createsLimit_and_hasLimit + CategoryTheory.preservesLimitOfShape_of_createsLimitsOfShape_and_hasLimitsOfShape /-- We construct a bijection between morphisms `L(A ⊗ B) ⟶ X` and morphisms `LA ⊗ LB ⟶ X`. This bijection has two key properties: @@ -231,8 +231,8 @@ noncomputable def bijection (A B : C) (X : D) : _ ≃ (i.obj ((reflector i).obj A) ⊗ i.obj ((reflector i).obj B) ⟶ i.obj X) := ((exp.adjunction _).homEquiv _ _).symm _ ≃ (i.obj ((reflector i).obj A ⊗ (reflector i).obj B) ⟶ i.obj X) := - haveI : Limits.PreservesLimits i := (reflectorAdjunction i).rightAdjointPreservesLimits - haveI := Limits.preservesSmallestLimitsOfPreservesLimits i + haveI : Limits.PreservesLimits i := (reflectorAdjunction i).rightAdjoint_preservesLimits + haveI := Limits.preservesSmallestLimits_of_preservesLimits i Iso.homCongr (prodComparisonIso _ _ _).symm (Iso.refl (i.obj X)) _ ≃ ((reflector i).obj A ⊗ (reflector i).obj B ⟶ X) := i.fullyFaithfulOfReflective.homEquiv.symm @@ -292,21 +292,21 @@ open Limits If a reflective subcategory is an exponential ideal, then the reflector preserves binary products. This is the converse of `exponentialIdeal_of_preserves_binary_products`. -/ -noncomputable def preservesBinaryProductsOfExponentialIdeal : +lemma preservesBinaryProducts_of_exponentialIdeal : PreservesLimitsOfShape (Discrete WalkingPair) (reflector i) where preservesLimit {K} := - letI := preservesLimitPairOfIsIsoProdComparison + letI := preservesLimit_pair_of_isIso_prodComparison (reflector i) (K.obj ⟨WalkingPair.left⟩) (K.obj ⟨WalkingPair.right⟩) - Limits.preservesLimitOfIsoDiagram _ (diagramIsoPair K).symm + Limits.preservesLimit_of_iso_diagram _ (diagramIsoPair K).symm /-- If a reflective subcategory is an exponential ideal, then the reflector preserves finite products. -/ -noncomputable def preservesFiniteProductsOfExponentialIdeal (J : Type) [Fintype J] : +lemma preservesFiniteProducts_of_exponentialIdeal (J : Type) [Fintype J] : PreservesLimitsOfShape (Discrete J) (reflector i) := by - letI := preservesBinaryProductsOfExponentialIdeal i - letI : PreservesLimitsOfShape _ (reflector i) := leftAdjointPreservesTerminalOfReflective.{0} i - apply preservesFiniteProductsOfPreservesBinaryAndTerminal (reflector i) J + letI := preservesBinaryProducts_of_exponentialIdeal i + letI : PreservesLimitsOfShape _ (reflector i) := leftAdjoint_preservesTerminal_of_reflective.{0} i + apply preservesFiniteProducts_of_preserves_binary_and_terminal (reflector i) J end diff --git a/Mathlib/CategoryTheory/Closed/Monoidal.lean b/Mathlib/CategoryTheory/Closed/Monoidal.lean index fec6606dada4e..4e97aea044bd9 100644 --- a/Mathlib/CategoryTheory/Closed/Monoidal.lean +++ b/Mathlib/CategoryTheory/Closed/Monoidal.lean @@ -115,7 +115,7 @@ end ihom open CategoryTheory.Limits instance : PreservesColimits (tensorLeft A) := - (ihom.adjunction A).leftAdjointPreservesColimits + (ihom.adjunction A).leftAdjoint_preservesColimits variable {A} diff --git a/Mathlib/CategoryTheory/Extensive.lean b/Mathlib/CategoryTheory/Extensive.lean index 4b4479017c601..87488121b4c56 100644 --- a/Mathlib/CategoryTheory/Extensive.lean +++ b/Mathlib/CategoryTheory/Extensive.lean @@ -146,18 +146,18 @@ variable [PreservesPullbacksOfInclusions F] {X Y Z : C} (f : Z ⟶ X ⨿ Y) noncomputable instance preservesPullbackInl' : PreservesLimit (cospan f coprod.inl) F := - preservesPullbackSymmetry _ _ _ + preservesPullback_symmetry _ _ _ noncomputable instance preservesPullbackInr' : PreservesLimit (cospan f coprod.inr) F := by - apply preservesLimitOfIsoDiagram (K₁ := cospan (f ≫ (coprod.braiding X Y).hom) coprod.inl) + apply preservesLimit_of_iso_diagram (K₁ := cospan (f ≫ (coprod.braiding X Y).hom) coprod.inl) apply cospanExt (Iso.refl _) (Iso.refl _) (coprod.braiding X Y).symm <;> simp noncomputable instance preservesPullbackInr : PreservesLimit (cospan coprod.inr f) F := - preservesPullbackSymmetry _ _ _ + preservesPullback_symmetry _ _ _ end PreservesPullbacksOfInclusions @@ -354,7 +354,7 @@ theorem finitaryExtensive_of_reflective [∀ X Y (f : X ⟶ Gl.obj Y), PreservesLimit (cospan (Gr.map f) (adj.unit.app Y)) Gl] [PreservesPullbacksOfInclusions Gl] : FinitaryExtensive D := by - have : PreservesColimitsOfSize Gl := adj.leftAdjointPreservesColimits + have : PreservesColimitsOfSize Gl := adj.leftAdjoint_preservesColimits constructor intros X Y c hc apply (IsVanKampenColimit.precompose_isIso_iff @@ -376,20 +376,18 @@ instance finitaryExtensive_functor [HasPullbacks C] [FinitaryExtensive C] : FinitaryExtensive (D ⥤ C) := haveI : HasFiniteCoproducts (D ⥤ C) := ⟨fun _ => Limits.functorCategoryHasColimitsOfShape⟩ ⟨fun c hc => isVanKampenColimit_of_evaluation _ c fun _ => - FinitaryExtensive.vanKampen _ <| PreservesColimit.preserves hc⟩ + FinitaryExtensive.vanKampen _ <| isColimitOfPreserves _ hc⟩ -noncomputable instance {C} [Category C] {D} [Category D] (F : C ⥤ D) {X Y Z : C} (f : X ⟶ Z) (g : Y ⟶ Z) [IsIso f] : PreservesLimit (cospan f g) F := have := hasPullback_of_left_iso f g - preservesLimitOfPreservesLimitCone (IsPullback.of_hasPullback f g).isLimit + preservesLimit_of_preserves_limit_cone (IsPullback.of_hasPullback f g).isLimit ((isLimitMapConePullbackConeEquiv _ pullback.condition).symm (IsPullback.of_vert_isIso ⟨by simp only [← F.map_comp, pullback.condition]⟩).isLimit) -noncomputable instance {C} [Category C] {D} [Category D] (F : C ⥤ D) {X Y Z : C} (f : X ⟶ Z) (g : Y ⟶ Z) [IsIso g] : PreservesLimit (cospan f g) F := - preservesPullbackSymmetry _ _ _ + preservesPullback_symmetry _ _ _ theorem finitaryExtensive_of_preserves_and_reflects (F : C ⥤ D) [FinitaryExtensive D] [HasFiniteCoproducts C] [HasPullbacksOfInclusions C] @@ -409,9 +407,9 @@ theorem finitaryExtensive_of_preserves_and_reflects_isomorphism (F : C ⥤ D) [F [HasFiniteCoproducts C] [HasPullbacks C] [PreservesLimitsOfShape WalkingCospan F] [PreservesColimitsOfShape (Discrete WalkingPair) F] [F.ReflectsIsomorphisms] : FinitaryExtensive C := by - haveI : ReflectsLimitsOfShape WalkingCospan F := reflectsLimitsOfShapeOfReflectsIsomorphisms + haveI : ReflectsLimitsOfShape WalkingCospan F := reflectsLimitsOfShape_of_reflectsIsomorphisms haveI : ReflectsColimitsOfShape (Discrete WalkingPair) F := - reflectsColimitsOfShapeOfReflectsIsomorphisms + reflectsColimitsOfShape_of_reflectsIsomorphisms exact finitaryExtensive_of_preserves_and_reflects F end Functor diff --git a/Mathlib/CategoryTheory/Functor/Flat.lean b/Mathlib/CategoryTheory/Functor/Flat.lean index 43e090be3c97b..2d0444d695f18 100644 --- a/Mathlib/CategoryTheory/Functor/Flat.lean +++ b/Mathlib/CategoryTheory/Functor/Flat.lean @@ -28,14 +28,14 @@ This definition is equivalent to left exact functors (functors that preserves fi * `flat_of_preservesFiniteLimits`: If `F : C ⥤ D` preserves finite limits and `C` has all finite limits, then `F` is flat. -* `preservesFiniteLimitsOfFlat`: If `F : C ⥤ D` is flat, then it preserves all finite limits. -* `preservesFiniteLimitsIffFlat`: If `C` has all finite limits, +* `preservesFiniteLimits_of_flat`: If `F : C ⥤ D` is flat, then it preserves all finite limits. +* `preservesFiniteLimits_iff_flat`: If `C` has all finite limits, then `F` is flat iff `F` is left_exact. -* `lanPreservesFiniteLimitsOfFlat`: If `F : C ⥤ D` is a flat functor between small categories, +* `lan_preservesFiniteLimits_of_flat`: If `F : C ⥤ D` is a flat functor between small categories, then the functor `Lan F.op` between presheaves of sets preserves all finite limits. * `flat_iff_lan_flat`: If `C`, `D` are small and `C` has all finite limits, then `F` is flat iff `Lan F.op : (Cᵒᵖ ⥤ Type*) ⥤ (Dᵒᵖ ⥤ Type*)` is flat. -* `preservesFiniteLimitsIffLanPreservesFiniteLimits`: If `C`, `D` are small and `C` has all +* `preservesFiniteLimits_iff_lanPreservesFiniteLimits`: If `C`, `D` are small and `C` has all finite limits, then `F` preserves finite limits iff `Lan F.op : (Cᵒᵖ ⥤ Type*) ⥤ (Dᵒᵖ ⥤ Type*)` does. @@ -157,7 +157,7 @@ theorem flat_of_preservesFiniteLimits [HasFiniteLimits C] (F : C ⥤ D) [Preserv theorem coflat_of_preservesFiniteColimits [HasFiniteColimits C] (F : C ⥤ D) [PreservesFiniteColimits F] : RepresentablyCoflat F := - let _ := preservesFiniteLimitsOp F + let _ := preservesFiniteLimits_op F (representablyFlat_op_iff _).1 (flat_of_preservesFiniteLimits _) namespace PreservesFiniteLimitsOfFlat @@ -239,12 +239,13 @@ theorem uniq {K : J ⥤ C} {c : Cone K} (hc : IsLimit c) (s : Cone (K ⋙ F)) end PreservesFiniteLimitsOfFlat /-- Representably flat functors preserve finite limits. -/ -noncomputable def preservesFiniteLimitsOfFlat (F : C ⥤ D) [RepresentablyFlat F] : +lemma preservesFiniteLimits_of_flat (F : C ⥤ D) [RepresentablyFlat F] : PreservesFiniteLimits F := by - apply preservesFiniteLimitsOfPreservesFiniteLimitsOfSize + apply preservesFiniteLimits_of_preservesFiniteLimitsOfSize intro J _ _; constructor intro K; constructor intro c hc + constructor exact { lift := PreservesFiniteLimitsOfFlat.lift F hc fac := PreservesFiniteLimitsOfFlat.fac F hc @@ -254,25 +255,23 @@ noncomputable def preservesFiniteLimitsOfFlat (F : C ⥤ D) [RepresentablyFlat F · exact PreservesFiniteLimitsOfFlat.fac F hc s } /-- Representably coflat functors preserve finite colimits. -/ -noncomputable def preservesFiniteColimitsOfCoflat (F : C ⥤ D) [RepresentablyCoflat F] : +lemma preservesFiniteColimits_of_coflat (F : C ⥤ D) [RepresentablyCoflat F] : PreservesFiniteColimits F := - letI _ := preservesFiniteLimitsOfFlat F.op - preservesFiniteColimitsOfOp _ + letI _ := preservesFiniteLimits_of_flat F.op + preservesFiniteColimits_of_op _ /-- If `C` is finitely complete, then `F : C ⥤ D` is representably flat iff it preserves finite limits. -/ -noncomputable def preservesFiniteLimitsIffFlat [HasFiniteLimits C] (F : C ⥤ D) : - RepresentablyFlat F ≃ PreservesFiniteLimits F := - equivOfSubsingletonOfSubsingleton - (fun _ => preservesFiniteLimitsOfFlat F) (fun _ => flat_of_preservesFiniteLimits F) +lemma preservesFiniteLimits_iff_flat [HasFiniteLimits C] (F : C ⥤ D) : + RepresentablyFlat F ↔ PreservesFiniteLimits F := + ⟨fun _ ↦ preservesFiniteLimits_of_flat F, fun _ ↦ flat_of_preservesFiniteLimits F⟩ /-- If `C` is finitely cocomplete, then `F : C ⥤ D` is representably coflat iff it preserves finite colmits. -/ -noncomputable def preservesFiniteColimitsIffCoflat [HasFiniteColimits C] (F : C ⥤ D) : - RepresentablyCoflat F ≃ PreservesFiniteColimits F := - equivOfSubsingletonOfSubsingleton - (fun _ => preservesFiniteColimitsOfCoflat F) (fun _ => coflat_of_preservesFiniteColimits F) +lemma preservesFiniteColimits_iff_coflat [HasFiniteColimits C] (F : C ⥤ D) : + RepresentablyCoflat F ↔ PreservesFiniteColimits F := + ⟨fun _ => preservesFiniteColimits_of_coflat F, fun _ => coflat_of_preservesFiniteColimits F⟩ end HasLimit @@ -311,15 +310,15 @@ variable [PreservesLimits (forget E)] /-- If `F : C ⥤ D` is a representably flat functor between small categories, then the functor `Lan F.op` that takes presheaves over `C` to presheaves over `D` preserves finite limits. -/ -noncomputable instance lanPreservesFiniteLimitsOfFlat (F : C ⥤ D) [RepresentablyFlat F] : +noncomputable instance lan_preservesFiniteLimits_of_flat (F : C ⥤ D) [RepresentablyFlat F] : PreservesFiniteLimits (F.op.lan : _ ⥤ Dᵒᵖ ⥤ E) := by - apply preservesFiniteLimitsOfPreservesFiniteLimitsOfSize.{u₁} + apply preservesFiniteLimits_of_preservesFiniteLimitsOfSize.{u₁} intro J _ _ - apply preservesLimitsOfShapeOfEvaluation (F.op.lan : (Cᵒᵖ ⥤ E) ⥤ Dᵒᵖ ⥤ E) J + apply preservesLimitsOfShape_of_evaluation (F.op.lan : (Cᵒᵖ ⥤ E) ⥤ Dᵒᵖ ⥤ E) J intro K haveI : IsFiltered (CostructuredArrow F.op K) := IsFiltered.of_equivalence (structuredArrowOpEquivalence F (unop K)) - exact preservesLimitsOfShapeOfNatIso (lanEvaluationIsoColim _ _ _).symm + exact preservesLimitsOfShape_of_natIso (lanEvaluationIsoColim _ _ _).symm instance lan_flat_of_flat (F : C ⥤ D) [RepresentablyFlat F] : RepresentablyFlat (F.op.lan : _ ⥤ Dᵒᵖ ⥤ E) := @@ -327,7 +326,7 @@ instance lan_flat_of_flat (F : C ⥤ D) [RepresentablyFlat F] : variable [HasFiniteLimits C] -noncomputable instance lanPreservesFiniteLimitsOfPreservesFiniteLimits (F : C ⥤ D) +instance lan_preservesFiniteLimits_of_preservesFiniteLimits (F : C ⥤ D) [PreservesFiniteLimits F] : PreservesFiniteLimits (F.op.lan : _ ⥤ Dᵒᵖ ⥤ E) := by haveI := flat_of_preservesFiniteLimits F infer_instance @@ -335,37 +334,20 @@ noncomputable instance lanPreservesFiniteLimitsOfPreservesFiniteLimits (F : C theorem flat_iff_lan_flat (F : C ⥤ D) : RepresentablyFlat F ↔ RepresentablyFlat (F.op.lan : _ ⥤ Dᵒᵖ ⥤ Type u₁) := ⟨fun _ => inferInstance, fun H => by - haveI := preservesFiniteLimitsOfFlat (F.op.lan : _ ⥤ Dᵒᵖ ⥤ Type u₁) + haveI := preservesFiniteLimits_of_flat (F.op.lan : _ ⥤ Dᵒᵖ ⥤ Type u₁) haveI : PreservesFiniteLimits F := by - apply preservesFiniteLimitsOfPreservesFiniteLimitsOfSize.{u₁} - intros; apply preservesLimitOfLanPreservesLimit + apply preservesFiniteLimits_of_preservesFiniteLimitsOfSize.{u₁} + intros; apply preservesLimit_of_lan_preservesLimit apply flat_of_preservesFiniteLimits⟩ /-- If `C` is finitely complete, then `F : C ⥤ D` preserves finite limits iff `Lan F.op : (Cᵒᵖ ⥤ Type*) ⥤ (Dᵒᵖ ⥤ Type*)` preserves finite limits. -/ -noncomputable def preservesFiniteLimitsIffLanPreservesFiniteLimits (F : C ⥤ D) : - PreservesFiniteLimits F ≃ PreservesFiniteLimits (F.op.lan : _ ⥤ Dᵒᵖ ⥤ Type u₁) where - toFun _ := inferInstance - invFun _ := by - apply preservesFiniteLimitsOfPreservesFiniteLimitsOfSize.{u₁} - intros; apply preservesLimitOfLanPreservesLimit - left_inv x := by - -- Porting note: `cases x` and an `unfold` not necessary in lean 4. - -- Remark : in mathlib3 we had `unfold preservesFiniteLimitsOfFlat` - -- but there was no `preservesFiniteLimitsOfFlat` in the goal! Experimentation - -- indicates that it was doing the same as `dsimp only` - dsimp only [preservesFiniteLimitsOfPreservesFiniteLimitsOfSize]; congr - -- Porting note: next line wasn't necessary in lean 3 - subsingleton - right_inv x := by - -- cases x; -- Porting note: not necessary in lean 4 - dsimp only [lanPreservesFiniteLimitsOfPreservesFiniteLimits, - lanPreservesFiniteLimitsOfFlat, - preservesFiniteLimitsOfPreservesFiniteLimitsOfSize] - congr - -- Porting note: next line wasn't necessary in lean 3 - subsingleton +lemma preservesFiniteLimits_iff_lan_preservesFiniteLimits (F : C ⥤ D) : + PreservesFiniteLimits F ↔ PreservesFiniteLimits (F.op.lan : _ ⥤ Dᵒᵖ ⥤ Type u₁) := + ⟨fun _ ↦ inferInstance, + fun _ ↦ preservesFiniteLimits_of_preservesFiniteLimitsOfSize.{u₁} _ + (fun _ _ _ ↦ preservesLimit_of_lan_preservesLimit _ _)⟩ end SmallCategory diff --git a/Mathlib/CategoryTheory/Galois/Action.lean b/Mathlib/CategoryTheory/Galois/Action.lean index fbf5fbd2f2b86..e45c24948c104 100644 --- a/Mathlib/CategoryTheory/Galois/Action.lean +++ b/Mathlib/CategoryTheory/Galois/Action.lean @@ -69,16 +69,16 @@ instance : Functor.ReflectsIsomorphisms (functorToAction F) where isIso_of_reflects_iso f F noncomputable instance : PreservesFiniteCoproducts (functorToAction F) := - ⟨fun J _ ↦ Action.preservesColimitsOfShapeOfPreserves (functorToAction F) + ⟨fun J _ ↦ Action.preservesColimitsOfShape_of_preserves (functorToAction F) (inferInstanceAs <| PreservesColimitsOfShape (Discrete J) F)⟩ noncomputable instance : PreservesFiniteProducts (functorToAction F) := - ⟨fun J _ ↦ Action.preservesLimitsOfShapeOfPreserves (functorToAction F) + ⟨fun J _ ↦ Action.preservesLimitsOfShape_of_preserves (functorToAction F) (inferInstanceAs <| PreservesLimitsOfShape (Discrete J) F)⟩ noncomputable instance (G : Type*) [Group G] [Finite G] : PreservesColimitsOfShape (SingleObj G) (functorToAction F) := - Action.preservesColimitsOfShapeOfPreserves _ <| + Action.preservesColimitsOfShape_of_preserves _ <| inferInstanceAs <| PreservesColimitsOfShape (SingleObj G) F instance : PreservesIsConnected (functorToAction F) := diff --git a/Mathlib/CategoryTheory/Galois/Basic.lean b/Mathlib/CategoryTheory/Galois/Basic.lean index 131a2e864a5bb..d6afd4303f42b 100644 --- a/Mathlib/CategoryTheory/Galois/Basic.lean +++ b/Mathlib/CategoryTheory/Galois/Basic.lean @@ -131,19 +131,19 @@ attribute [instance] preservesTerminalObjects preservesPullbacks preservesEpis preservesFiniteCoproducts reflectsIsos preservesQuotientsByFiniteGroups noncomputable instance : ReflectsLimitsOfShape (Discrete PEmpty.{1}) F := - reflectsLimitsOfShapeOfReflectsIsomorphisms + reflectsLimitsOfShape_of_reflectsIsomorphisms noncomputable instance : ReflectsColimitsOfShape (Discrete PEmpty.{1}) F := - reflectsColimitsOfShapeOfReflectsIsomorphisms + reflectsColimitsOfShape_of_reflectsIsomorphisms noncomputable instance : PreservesFiniteLimits F := - preservesFiniteLimitsOfPreservesTerminalAndPullbacks F + preservesFiniteLimits_of_preservesTerminal_and_pullbacks F /-- Fiber functors preserve quotients by finite groups in arbitrary universes. -/ -noncomputable instance {G : Type*} [Group G] [Finite G] : +instance {G : Type*} [Group G] [Finite G] : PreservesColimitsOfShape (SingleObj G) F := by choose G' hg hf he using Finite.exists_type_univ_nonempty_mulEquiv G - exact Limits.preservesColimitsOfShapeOfEquiv he.some.toSingleObjEquiv.symm F + exact Limits.preservesColimitsOfShape_of_equiv he.some.toSingleObjEquiv.symm F /-- Fiber functors reflect monomorphisms. -/ instance : ReflectsMonomorphisms F := ReflectsMonomorphisms.mk <| by @@ -170,9 +170,9 @@ section /-- If `F` is a fiber functor and `E` is an equivalence between categories of finite types, then `F ⋙ E` is again a fiber functor. -/ -noncomputable def compRight (E : FintypeCat.{w} ⥤ FintypeCat.{t}) [E.IsEquivalence] : +lemma comp_right (E : FintypeCat.{w} ⥤ FintypeCat.{t}) [E.IsEquivalence] : FiberFunctor (F ⋙ E) where - preservesQuotientsByFiniteGroups _ := compPreservesColimitsOfShape F E + preservesQuotientsByFiniteGroups _ := comp_preservesColimitsOfShape F E end diff --git a/Mathlib/CategoryTheory/Galois/Examples.lean b/Mathlib/CategoryTheory/Galois/Examples.lean index f52783a0ae08a..046ec34c9ba91 100644 --- a/Mathlib/CategoryTheory/Galois/Examples.lean +++ b/Mathlib/CategoryTheory/Galois/Examples.lean @@ -75,7 +75,7 @@ instance [Finite G] : HasColimitsOfShape (SingleObj G) FintypeCat.{w} := by noncomputable instance : PreservesFiniteLimits (forget (Action FintypeCat (MonCat.of G))) := by show PreservesFiniteLimits (Action.forget FintypeCat _ ⋙ FintypeCat.incl) - apply compPreservesFiniteLimits + apply comp_preservesFiniteLimits /-- The category of finite `G`-sets is a `PreGaloisCategory`. -/ instance : PreGaloisCategory (Action FintypeCat (MonCat.of G)) where diff --git a/Mathlib/CategoryTheory/Galois/GaloisObjects.lean b/Mathlib/CategoryTheory/Galois/GaloisObjects.lean index 9c838d2ad07d4..89a2b0fb5a9ac 100644 --- a/Mathlib/CategoryTheory/Galois/GaloisObjects.lean +++ b/Mathlib/CategoryTheory/Galois/GaloisObjects.lean @@ -37,7 +37,7 @@ open Limits Functor noncomputable instance {G : Type v} [Group G] [Finite G] : PreservesColimitsOfShape (SingleObj G) FintypeCat.incl.{w} := by choose G' hg hf e using Finite.exists_type_univ_nonempty_mulEquiv G - exact Limits.preservesColimitsOfShapeOfEquiv (Classical.choice e).toSingleObjEquiv.symm _ + exact Limits.preservesColimitsOfShape_of_equiv (Classical.choice e).toSingleObjEquiv.symm _ /-- A connected object `X` of `C` is Galois if the quotient `X / Aut X` is terminal. -/ class IsGalois {C : Type u₁} [Category.{u₂, u₁} C] [GaloisCategory C] (X : C) diff --git a/Mathlib/CategoryTheory/Galois/Prorepresentability.lean b/Mathlib/CategoryTheory/Galois/Prorepresentability.lean index fba4b027d12de..554c363a78c89 100644 --- a/Mathlib/CategoryTheory/Galois/Prorepresentability.lean +++ b/Mathlib/CategoryTheory/Galois/Prorepresentability.lean @@ -449,7 +449,7 @@ instance FiberFunctor.isPretransitive_of_isConnected (X : C) [IsConnected X] : MulAction.IsPretransitive (Aut F) (F.obj X) where exists_smul_eq x y := by let F' : C ⥤ FintypeCat.{u₂} := F ⋙ FintypeCat.uSwitch.{w, u₂} - letI : FiberFunctor F' := FiberFunctor.compRight _ + letI : FiberFunctor F' := FiberFunctor.comp_right _ let e (Y : C) : F'.obj Y ≃ F.obj Y := (F.obj Y).uSwitchEquiv set x' : F'.obj X := (e X).symm x with hx' set y' : F'.obj X := (e X).symm y with hy' diff --git a/Mathlib/CategoryTheory/GlueData.lean b/Mathlib/CategoryTheory/GlueData.lean index 69eb61fe1f99e..936283e96bec5 100644 --- a/Mathlib/CategoryTheory/GlueData.lean +++ b/Mathlib/CategoryTheory/GlueData.lean @@ -299,7 +299,7 @@ end variable [HasMulticoequalizer D.diagram] [PreservesColimit D.diagram.multispan F] theorem hasColimit_multispan_comp : HasColimit (D.diagram.multispan ⋙ F) := - ⟨⟨⟨_, PreservesColimit.preserves (colimit.isColimit _)⟩⟩⟩ + ⟨⟨⟨_, isColimitOfPreserves _ (colimit.isColimit _)⟩⟩⟩ attribute [local instance] hasColimit_multispan_comp diff --git a/Mathlib/CategoryTheory/GradedObject/Monoidal.lean b/Mathlib/CategoryTheory/GradedObject/Monoidal.lean index b005a034c6b49..34da9a269a20b 100644 --- a/Mathlib/CategoryTheory/GradedObject/Monoidal.lean +++ b/Mathlib/CategoryTheory/GradedObject/Monoidal.lean @@ -546,6 +546,7 @@ lemma triangle : tensorHom (rightUnitor X₁).hom (𝟙 X₃) := by convert mapBifunctor_triangle (curriedAssociatorNatIso C) (𝟙_ C) (rightUnitorNatIso C) (leftUnitorNatIso C) (triangleIndexData I) X₁ X₃ (by simp) + all_goals assumption end Triangle diff --git a/Mathlib/CategoryTheory/Limits/Comma.lean b/Mathlib/CategoryTheory/Limits/Comma.lean index f67006d4500bf..93476394301e1 100644 --- a/Mathlib/CategoryTheory/Limits/Comma.lean +++ b/Mathlib/CategoryTheory/Limits/Comma.lean @@ -48,7 +48,7 @@ limit cone for `F ⋙ snd L R : J ⥤ R` we can build a cone for `F` which will cone. -/ @[simps] -def coneOfPreserves [PreservesLimit (F ⋙ snd L R) R] (c₁ : Cone (F ⋙ fst L R)) +noncomputable def coneOfPreserves [PreservesLimit (F ⋙ snd L R) R] (c₁ : Cone (F ⋙ fst L R)) {c₂ : Cone (F ⋙ snd L R)} (t₂ : IsLimit c₂) : Cone F where pt := { left := c₁.pt @@ -66,7 +66,7 @@ def coneOfPreserves [PreservesLimit (F ⋙ snd L R) R] (c₁ : Cone (F ⋙ fst L /-- Provided that `R` preserves the appropriate limit, then the cone in `coneOfPreserves` is a limit. -/ -def coneOfPreservesIsLimit [PreservesLimit (F ⋙ snd L R) R] {c₁ : Cone (F ⋙ fst L R)} +noncomputable def coneOfPreservesIsLimit [PreservesLimit (F ⋙ snd L R) R] {c₁ : Cone (F ⋙ fst L R)} (t₁ : IsLimit c₁) {c₂ : Cone (F ⋙ snd L R)} (t₂ : IsLimit c₂) : IsLimit (coneOfPreserves F c₁ t₂) where lift s := @@ -95,7 +95,7 @@ a cocone for `F ⋙ snd L R : J ⥤ R` we can build a cocone for `F` which will colimit cocone. -/ @[simps] -def coconeOfPreserves [PreservesColimit (F ⋙ fst L R) L] {c₁ : Cocone (F ⋙ fst L R)} +noncomputable def coconeOfPreserves [PreservesColimit (F ⋙ fst L R) L] {c₁ : Cocone (F ⋙ fst L R)} (t₁ : IsColimit c₁) (c₂ : Cocone (F ⋙ snd L R)) : Cocone F where pt := { left := c₁.pt @@ -113,7 +113,8 @@ def coconeOfPreserves [PreservesColimit (F ⋙ fst L R) L] {c₁ : Cocone (F ⋙ /-- Provided that `L` preserves the appropriate colimit, then the cocone in `coconeOfPreserves` is a colimit. -/ -def coconeOfPreservesIsColimit [PreservesColimit (F ⋙ fst L R) L] {c₁ : Cocone (F ⋙ fst L R)} +noncomputable def coconeOfPreservesIsColimit [PreservesColimit (F ⋙ fst L R) L] + {c₁ : Cocone (F ⋙ fst L R)} (t₁ : IsColimit c₁) {c₂ : Cocone (F ⋙ snd L R)} (t₂ : IsColimit c₂) : IsColimit (coconeOfPreserves F t₁ c₂) where desc s := diff --git a/Mathlib/CategoryTheory/Limits/ConeCategory.lean b/Mathlib/CategoryTheory/Limits/ConeCategory.lean index f7427f3b61376..f15da439d1804 100644 --- a/Mathlib/CategoryTheory/Limits/ConeCategory.lean +++ b/Mathlib/CategoryTheory/Limits/ConeCategory.lean @@ -181,12 +181,12 @@ theorem IsTerminal.from_eq_liftConeMorphism {F : J ⥤ C} {c : Cone F} (hc : IsT (IsLimit.liftConeMorphism_eq_isTerminal_from (c.isLimitEquivIsTerminal.symm hc) s).symm /-- If `G : Cone F ⥤ Cone F'` preserves terminal objects, it preserves limit cones. -/ -def IsLimit.ofPreservesConeTerminal {F : J ⥤ C} {F' : K ⥤ D} (G : Cone F ⥤ Cone F') +noncomputable def IsLimit.ofPreservesConeTerminal {F : J ⥤ C} {F' : K ⥤ D} (G : Cone F ⥤ Cone F') [PreservesLimit (Functor.empty.{0} _) G] {c : Cone F} (hc : IsLimit c) : IsLimit (G.obj c) := (Cone.isLimitEquivIsTerminal _).symm <| (Cone.isLimitEquivIsTerminal _ hc).isTerminalObj _ _ /-- If `G : Cone F ⥤ Cone F'` reflects terminal objects, it reflects limit cones. -/ -def IsLimit.ofReflectsConeTerminal {F : J ⥤ C} {F' : K ⥤ D} (G : Cone F ⥤ Cone F') +noncomputable def IsLimit.ofReflectsConeTerminal {F : J ⥤ C} {F' : K ⥤ D} (G : Cone F ⥤ Cone F') [ReflectsLimit (Functor.empty.{0} _) G] {c : Cone F} (hc : IsLimit (G.obj c)) : IsLimit c := (Cone.isLimitEquivIsTerminal _).symm <| (Cone.isLimitEquivIsTerminal _ hc).isTerminalOfObj _ _ @@ -341,13 +341,15 @@ theorem IsInitial.to_eq_descCoconeMorphism {F : J ⥤ C} {c : Cocone F} (hc : Is (IsColimit.descCoconeMorphism_eq_isInitial_to (c.isColimitEquivIsInitial.symm hc) s).symm /-- If `G : Cocone F ⥤ Cocone F'` preserves initial objects, it preserves colimit cocones. -/ -def IsColimit.ofPreservesCoconeInitial {F : J ⥤ C} {F' : K ⥤ D} (G : Cocone F ⥤ Cocone F') +noncomputable def IsColimit.ofPreservesCoconeInitial {F : J ⥤ C} {F' : K ⥤ D} + (G : Cocone F ⥤ Cocone F') [PreservesColimit (Functor.empty.{0} _) G] {c : Cocone F} (hc : IsColimit c) : IsColimit (G.obj c) := (Cocone.isColimitEquivIsInitial _).symm <| (Cocone.isColimitEquivIsInitial _ hc).isInitialObj _ _ /-- If `G : Cocone F ⥤ Cocone F'` reflects initial objects, it reflects colimit cocones. -/ -def IsColimit.ofReflectsCoconeInitial {F : J ⥤ C} {F' : K ⥤ D} (G : Cocone F ⥤ Cocone F') +noncomputable def IsColimit.ofReflectsCoconeInitial {F : J ⥤ C} {F' : K ⥤ D} + (G : Cocone F ⥤ Cocone F') [ReflectsColimit (Functor.empty.{0} _) G] {c : Cocone F} (hc : IsColimit (G.obj c)) : IsColimit c := (Cocone.isColimitEquivIsInitial _).symm <| diff --git a/Mathlib/CategoryTheory/Limits/Connected.lean b/Mathlib/CategoryTheory/Limits/Connected.lean index 712b12fbf1148..0d9777ebc99b8 100644 --- a/Mathlib/CategoryTheory/Limits/Connected.lean +++ b/Mathlib/CategoryTheory/Limits/Connected.lean @@ -87,12 +87,11 @@ Note that this functor does not preserve the two most obvious disconnected limit `(X × -)` does not preserve products or terminal object, eg `(X ⨯ A) ⨯ (X ⨯ B)` is not isomorphic to `X ⨯ (A ⨯ B)` and `X ⨯ 1` is not isomorphic to `1`. -/ -noncomputable def prodPreservesConnectedLimits [IsConnected J] (X : C) : +lemma prod_preservesConnectedLimits [IsConnected J] (X : C) : PreservesLimitsOfShape J (prod.functor.obj X) where preservesLimit {K} := - { - preserves := fun {c} l => - { lift := fun s => + { preserves := fun {c} l => ⟨{ + lift := fun s => prod.lift (s.π.app (Classical.arbitrary _) ≫ Limits.prod.fst) (l.lift (forgetCone s)) fac := fun s j => by apply Limits.prod.hom_ext @@ -106,6 +105,6 @@ noncomputable def prodPreservesConnectedLimits [IsConnected J] (X : C) : · rw [limit.lift_π] apply l.uniq (forgetCone s) intro j - simp [← L j] } } + simp [← L j] }⟩ } end CategoryTheory diff --git a/Mathlib/CategoryTheory/Limits/Constructions/BinaryProducts.lean b/Mathlib/CategoryTheory/Limits/Constructions/BinaryProducts.lean index 3278a8a966db7..4ad7a58d9ac61 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/BinaryProducts.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/BinaryProducts.lean @@ -93,11 +93,11 @@ theorem hasBinaryProducts_of_hasTerminal_and_pullbacks [HasTerminal C] [HasPullb variable {C} /-- A functor that preserves terminal objects and pullbacks preserves binary products. -/ -noncomputable def preservesBinaryProductsOfPreservesTerminalAndPullbacks [HasTerminal C] +lemma preservesBinaryProducts_of_preservesTerminal_and_pullbacks [HasTerminal C] [HasPullbacks C] [PreservesLimitsOfShape (Discrete.{0} PEmpty) F] [PreservesLimitsOfShape WalkingCospan F] : PreservesLimitsOfShape (Discrete WalkingPair) F := ⟨fun {K} => - preservesLimitOfPreservesLimitCone (limitConeOfTerminalAndPullbacks K).2 + preservesLimit_of_preserves_limit_cone (limitConeOfTerminalAndPullbacks K).2 (by apply isBinaryProductOfIsTerminalIsPullback _ _ (isLimitOfHasTerminalOfPreservesLimit F) @@ -199,11 +199,11 @@ theorem hasBinaryCoproducts_of_hasInitial_and_pushouts [HasInitial C] [HasPushou variable {C} /-- A functor that preserves initial objects and pushouts preserves binary coproducts. -/ -noncomputable def preservesBinaryCoproductsOfPreservesInitialAndPushouts [HasInitial C] +lemma preservesBinaryCoproducts_of_preservesInitial_and_pushouts [HasInitial C] [HasPushouts C] [PreservesColimitsOfShape (Discrete.{0} PEmpty) F] [PreservesColimitsOfShape WalkingSpan F] : PreservesColimitsOfShape (Discrete WalkingPair) F := ⟨fun {K} => - preservesColimitOfPreservesColimitCocone (colimitCoconeOfInitialAndPushouts K).2 (by + preservesColimit_of_preserves_colimit_cocone (colimitCoconeOfInitialAndPushouts K).2 (by apply isBinaryCoproductOfIsInitialIsPushout _ _ (isColimitOfHasInitialOfPreservesColimit F) diff --git a/Mathlib/CategoryTheory/Limits/Constructions/Equalizers.lean b/Mathlib/CategoryTheory/Limits/Constructions/Equalizers.lean index aa87a67935f4c..73d519ebd5219 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/Equalizers.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/Equalizers.lean @@ -86,11 +86,12 @@ theorem hasEqualizers_of_hasPullbacks_and_binary_products [HasBinaryProducts C] attribute [local instance] hasPullback_of_preservesPullback /-- A functor that preserves pullbacks and binary products also presrves equalizers. -/ -def preservesEqualizersOfPreservesPullbacksAndBinaryProducts [HasBinaryProducts C] [HasPullbacks C] +lemma preservesEqualizers_of_preservesPullbacks_and_binaryProducts + [HasBinaryProducts C] [HasPullbacks C] [PreservesLimitsOfShape (Discrete WalkingPair) G] [PreservesLimitsOfShape WalkingCospan G] : PreservesLimitsOfShape WalkingParallelPair G := ⟨fun {K} => - preservesLimitOfPreservesLimitCone (equalizerConeIsLimit K) <| + preservesLimit_of_preserves_limit_cone (equalizerConeIsLimit K) <| { lift := fun c => by refine pullback.lift ?_ ?_ ?_ ≫ (PreservesPullback.iso _ _ _ ).inv · exact c.π.app WalkingParallelPair.zero @@ -180,11 +181,11 @@ theorem hasCoequalizers_of_hasPushouts_and_binary_coproducts [HasBinaryCoproduct attribute [local instance] hasPushout_of_preservesPushout /-- A functor that preserves pushouts and binary coproducts also presrves coequalizers. -/ -def preservesCoequalizersOfPreservesPushoutsAndBinaryCoproducts [HasBinaryCoproducts C] +lemma preservesCoequalizers_of_preservesPushouts_and_binaryCoproducts [HasBinaryCoproducts C] [HasPushouts C] [PreservesColimitsOfShape (Discrete WalkingPair) G] [PreservesColimitsOfShape WalkingSpan G] : PreservesColimitsOfShape WalkingParallelPair G := ⟨fun {K} => - preservesColimitOfPreservesColimitCocone (coequalizerCoconeIsColimit K) <| + preservesColimit_of_preserves_colimit_cocone (coequalizerCoconeIsColimit K) <| { desc := fun c => by refine (PreservesPushout.iso _ _ _).inv ≫ pushout.desc ?_ ?_ ?_ · exact c.ι.app WalkingParallelPair.one diff --git a/Mathlib/CategoryTheory/Limits/Constructions/FiniteProductsOfBinaryProducts.lean b/Mathlib/CategoryTheory/Limits/Constructions/FiniteProductsOfBinaryProducts.lean index ffb743849fa51..629aa364f506a 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/FiniteProductsOfBinaryProducts.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/FiniteProductsOfBinaryProducts.lean @@ -117,17 +117,17 @@ variable [HasFiniteProducts.{v} C] /-- If `F` preserves the terminal object and binary products, then it preserves products indexed by `Fin n` for any `n`. -/ -noncomputable def preservesFinOfPreservesBinaryAndTerminal : +lemma preservesFinOfPreservesBinaryAndTerminal : ∀ (n : ℕ) (f : Fin n → C), PreservesLimit (Discrete.functor f) F | 0 => fun f => by letI : PreservesLimitsOfShape (Discrete (Fin 0)) F := - preservesLimitsOfShapeOfEquiv.{0, 0} (Discrete.equivalence finZeroEquiv'.symm) _ + preservesLimitsOfShape_of_equiv.{0, 0} (Discrete.equivalence finZeroEquiv'.symm) _ infer_instance | n + 1 => by haveI := preservesFinOfPreservesBinaryAndTerminal n intro f apply - preservesLimitOfPreservesLimitCone + preservesLimit_of_preserves_limit_cone (extendFanIsLimit f (limit.isLimit _) (limit.isLimit _)) _ apply (isLimitMapConeFanMkEquiv _ _ _).symm _ let this := @@ -148,20 +148,20 @@ noncomputable def preservesFinOfPreservesBinaryAndTerminal : /-- If `F` preserves the terminal object and binary products, then it preserves limits of shape `Discrete (Fin n)`. -/ -def preservesShapeFinOfPreservesBinaryAndTerminal (n : ℕ) : +lemma preservesShape_fin_of_preserves_binary_and_terminal (n : ℕ) : PreservesLimitsOfShape (Discrete (Fin n)) F where preservesLimit {K} := by let that : (Discrete.functor fun n => K.obj ⟨n⟩) ≅ K := Discrete.natIso fun ⟨i⟩ => Iso.refl _ haveI := preservesFinOfPreservesBinaryAndTerminal F n fun n => K.obj ⟨n⟩ - apply preservesLimitOfIsoDiagram F that + apply preservesLimit_of_iso_diagram F that /-- If `F` preserves the terminal object and binary products then it preserves finite products. -/ -def preservesFiniteProductsOfPreservesBinaryAndTerminal (J : Type*) [Fintype J] : +lemma preservesFiniteProducts_of_preserves_binary_and_terminal (J : Type*) [Fintype J] : PreservesLimitsOfShape (Discrete J) F := by classical let e := Fintype.equivFin J - haveI := preservesShapeFinOfPreservesBinaryAndTerminal F (Fintype.card J) - apply preservesLimitsOfShapeOfEquiv (Discrete.equivalence e).symm + haveI := preservesShape_fin_of_preserves_binary_and_terminal F (Fintype.card J) + apply preservesLimitsOfShape_of_equiv (Discrete.equivalence e).symm end Preserves @@ -249,17 +249,17 @@ variable [HasFiniteCoproducts.{v} C] /-- If `F` preserves the initial object and binary coproducts, then it preserves products indexed by `Fin n` for any `n`. -/ -noncomputable def preservesFinOfPreservesBinaryAndInitial : +lemma preserves_fin_of_preserves_binary_and_initial : ∀ (n : ℕ) (f : Fin n → C), PreservesColimit (Discrete.functor f) F | 0 => fun f => by letI : PreservesColimitsOfShape (Discrete (Fin 0)) F := - preservesColimitsOfShapeOfEquiv.{0, 0} (Discrete.equivalence finZeroEquiv'.symm) _ + preservesColimitsOfShape_of_equiv.{0, 0} (Discrete.equivalence finZeroEquiv'.symm) _ infer_instance | n + 1 => by - haveI := preservesFinOfPreservesBinaryAndInitial n + haveI := preserves_fin_of_preserves_binary_and_initial n intro f apply - preservesColimitOfPreservesColimitCocone + preservesColimit_of_preserves_colimit_cocone (extendCofanIsColimit f (colimit.isColimit _) (colimit.isColimit _)) _ apply (isColimitMapCoconeCofanMkEquiv _ _ _).symm _ let this := @@ -279,20 +279,20 @@ noncomputable def preservesFinOfPreservesBinaryAndInitial : /-- If `F` preserves the initial object and binary coproducts, then it preserves colimits of shape `Discrete (Fin n)`. -/ -def preservesShapeFinOfPreservesBinaryAndInitial (n : ℕ) : +lemma preservesShape_fin_of_preserves_binary_and_initial (n : ℕ) : PreservesColimitsOfShape (Discrete (Fin n)) F where preservesColimit {K} := by let that : (Discrete.functor fun n => K.obj ⟨n⟩) ≅ K := Discrete.natIso fun ⟨i⟩ => Iso.refl _ - haveI := preservesFinOfPreservesBinaryAndInitial F n fun n => K.obj ⟨n⟩ - apply preservesColimitOfIsoDiagram F that + haveI := preserves_fin_of_preserves_binary_and_initial F n fun n => K.obj ⟨n⟩ + apply preservesColimit_of_iso_diagram F that /-- If `F` preserves the initial object and binary coproducts then it preserves finite products. -/ -def preservesFiniteCoproductsOfPreservesBinaryAndInitial (J : Type*) [Fintype J] : +lemma preservesFiniteCoproductsOfPreservesBinaryAndInitial (J : Type*) [Fintype J] : PreservesColimitsOfShape (Discrete J) F := by classical let e := Fintype.equivFin J - haveI := preservesShapeFinOfPreservesBinaryAndInitial F (Fintype.card J) - apply preservesColimitsOfShapeOfEquiv (Discrete.equivalence e).symm + haveI := preservesShape_fin_of_preserves_binary_and_initial F (Fintype.card J) + apply preservesColimitsOfShape_of_equiv (Discrete.equivalence e).symm end Preserves diff --git a/Mathlib/CategoryTheory/Limits/Constructions/LimitsOfProductsAndEqualizers.lean b/Mathlib/CategoryTheory/Limits/Constructions/LimitsOfProductsAndEqualizers.lean index bd52c4c1ea3b7..353032099523c 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/LimitsOfProductsAndEqualizers.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/LimitsOfProductsAndEqualizers.lean @@ -156,7 +156,8 @@ variable (G : C ⥤ D) [PreservesLimitsOfShape WalkingParallelPair G] [PreservesLimitsOfShape (Discrete.{w} (Σp : J × J, p.1 ⟶ p.2)) G] /-- If a functor preserves equalizers and the appropriate products, it preserves limits. -/ -noncomputable def preservesLimitOfPreservesEqualizersAndProduct : PreservesLimitsOfShape J G where +lemma preservesLimit_of_preservesEqualizers_and_product : + PreservesLimitsOfShape J G where preservesLimit {K} := by let P := ∏ᶜ K.obj let Q := ∏ᶜ fun f : Σp : J × J, p.fst ⟶ p.snd => K.obj f.1.2 @@ -164,8 +165,7 @@ noncomputable def preservesLimitOfPreservesEqualizersAndProduct : PreservesLimit let t : P ⟶ Q := Pi.lift fun f => limit.π (Discrete.functor K.obj) ⟨f.1.2⟩ let I := equalizer s t let i : I ⟶ P := equalizer.ι s t - apply - preservesLimitOfPreservesLimitCone + apply preservesLimit_of_preserves_limit_cone (buildIsLimit s t (by simp [s]) (by simp [t]) (limit.isLimit _) (limit.isLimit _) (limit.isLimit _)) apply IsLimit.ofIsoLimit (buildIsLimit _ _ _ _ _ _ _) _ @@ -202,24 +202,24 @@ with type Fintype J does not have forward dependencies, type class resolution ca use this kind of local instance because it will not be able to infer a value for this parameter." Factored out this as new class in `CategoryTheory.Limits.Preserves.Finite` -/ /-- If G preserves equalizers and finite products, it preserves finite limits. -/ -noncomputable def preservesFiniteLimitsOfPreservesEqualizersAndFiniteProducts [HasEqualizers C] +lemma preservesFiniteLimits_of_preservesEqualizers_and_finiteProducts [HasEqualizers C] [HasFiniteProducts C] (G : C ⥤ D) [PreservesLimitsOfShape WalkingParallelPair G] [PreservesFiniteProducts G] : PreservesFiniteLimits G where preservesFiniteLimits := by intro J sJ fJ haveI : Fintype J := inferInstance haveI : Fintype ((p : J × J) × (p.fst ⟶ p.snd)) := inferInstance - apply @preservesLimitOfPreservesEqualizersAndProduct _ _ _ sJ _ _ ?_ ?_ _ G _ ?_ ?_ + apply @preservesLimit_of_preservesEqualizers_and_product _ _ _ sJ _ _ ?_ ?_ _ G _ ?_ ?_ · apply hasLimitsOfShape_discrete _ _ · apply hasLimitsOfShape_discrete _ · apply PreservesFiniteProducts.preserves _ · apply PreservesFiniteProducts.preserves _ /-- If G preserves equalizers and products, it preserves all limits. -/ -noncomputable def preservesLimitsOfPreservesEqualizersAndProducts [HasEqualizers C] +lemma preservesLimits_of_preservesEqualizers_and_products [HasEqualizers C] [HasProducts.{w} C] (G : C ⥤ D) [PreservesLimitsOfShape WalkingParallelPair G] [∀ J, PreservesLimitsOfShape (Discrete.{w} J) G] : PreservesLimitsOfSize.{w, w} G where - preservesLimitsOfShape := preservesLimitOfPreservesEqualizersAndProduct G + preservesLimitsOfShape := preservesLimit_of_preservesEqualizers_and_product G theorem hasFiniteLimits_of_hasTerminal_and_pullbacks [HasTerminal C] [HasPullbacks C] : HasFiniteLimits C := @@ -230,18 +230,18 @@ theorem hasFiniteLimits_of_hasTerminal_and_pullbacks [HasTerminal C] [HasPullbac (hasBinaryProducts_of_hasTerminal_and_pullbacks C) inferInstance) /-- If G preserves terminal objects and pullbacks, it preserves all finite limits. -/ -noncomputable def preservesFiniteLimitsOfPreservesTerminalAndPullbacks [HasTerminal C] +lemma preservesFiniteLimits_of_preservesTerminal_and_pullbacks [HasTerminal C] [HasPullbacks C] (G : C ⥤ D) [PreservesLimitsOfShape (Discrete.{0} PEmpty) G] [PreservesLimitsOfShape WalkingCospan G] : PreservesFiniteLimits G := by haveI : HasFiniteLimits C := hasFiniteLimits_of_hasTerminal_and_pullbacks haveI : PreservesLimitsOfShape (Discrete WalkingPair) G := - preservesBinaryProductsOfPreservesTerminalAndPullbacks G + preservesBinaryProducts_of_preservesTerminal_and_pullbacks G haveI : PreservesLimitsOfShape WalkingParallelPair G := - preservesEqualizersOfPreservesPullbacksAndBinaryProducts G + preservesEqualizers_of_preservesPullbacks_and_binaryProducts G apply - @preservesFiniteLimitsOfPreservesEqualizersAndFiniteProducts _ _ _ _ _ _ G _ ?_ + @preservesFiniteLimits_of_preservesEqualizers_and_finiteProducts _ _ _ _ _ _ G _ ?_ apply PreservesFiniteProducts.mk - apply preservesFiniteProductsOfPreservesBinaryAndTerminal G + apply preservesFiniteProducts_of_preserves_binary_and_terminal G /-! We now dualize the above constructions, resorting to copy-paste. @@ -372,7 +372,7 @@ variable (G : C ⥤ D) [PreservesColimitsOfShape WalkingParallelPair G] [PreservesColimitsOfShape (Discrete.{w} (Σp : J × J, p.1 ⟶ p.2)) G] /-- If a functor preserves coequalizers and the appropriate coproducts, it preserves colimits. -/ -noncomputable def preservesColimitOfPreservesCoequalizersAndCoproduct : +lemma preservesColimit_of_preservesCoequalizers_and_coproduct : PreservesColimitsOfShape J G where preservesColimit {K} := by let P := ∐ K.obj @@ -381,8 +381,7 @@ noncomputable def preservesColimitOfPreservesCoequalizersAndCoproduct : let t : Q ⟶ P := Sigma.desc fun f => colimit.ι (Discrete.functor K.obj) ⟨f.1.1⟩ let I := coequalizer s t let i : P ⟶ I := coequalizer.π s t - apply - preservesColimitOfPreservesColimitCocone + apply preservesColimit_of_preserves_colimit_cocone (buildIsColimit s t (by simp [s]) (by simp [t]) (colimit.isColimit _) (colimit.isColimit _) (colimit.isColimit _)) apply IsColimit.ofIsoColimit (buildIsColimit _ _ _ _ _ _ _) _ @@ -422,7 +421,7 @@ with type Fintype J does not have forward dependencies, type class resolution ca this kind of local instance because it will not be able to infer a value for this parameter." Factored out this as new class in `CategoryTheory.Limits.Preserves.Finite` -/ /-- If G preserves coequalizers and finite coproducts, it preserves finite colimits. -/ -noncomputable def preservesFiniteColimitsOfPreservesCoequalizersAndFiniteCoproducts +lemma preservesFiniteColimits_of_preservesCoequalizers_and_finiteCoproducts [HasCoequalizers C] [HasFiniteCoproducts C] (G : C ⥤ D) [PreservesColimitsOfShape WalkingParallelPair G] [PreservesFiniteCoproducts G] : PreservesFiniteColimits G where @@ -430,17 +429,17 @@ noncomputable def preservesFiniteColimitsOfPreservesCoequalizersAndFiniteCoprodu intro J sJ fJ haveI : Fintype J := inferInstance haveI : Fintype ((p : J × J) × (p.fst ⟶ p.snd)) := inferInstance - apply @preservesColimitOfPreservesCoequalizersAndCoproduct _ _ _ sJ _ _ ?_ ?_ _ G _ ?_ ?_ + apply @preservesColimit_of_preservesCoequalizers_and_coproduct _ _ _ sJ _ _ ?_ ?_ _ G _ ?_ ?_ · apply hasColimitsOfShape_discrete _ _ · apply hasColimitsOfShape_discrete _ · apply PreservesFiniteCoproducts.preserves _ · apply PreservesFiniteCoproducts.preserves _ /-- If G preserves coequalizers and coproducts, it preserves all colimits. -/ -noncomputable def preservesColimitsOfPreservesCoequalizersAndCoproducts [HasCoequalizers C] +lemma preservesColimits_of_preservesCoequalizers_and_coproducts [HasCoequalizers C] [HasCoproducts.{w} C] (G : C ⥤ D) [PreservesColimitsOfShape WalkingParallelPair G] - [∀ J, PreservesColimitsOfShape (Discrete.{w} J) G] : PreservesColimitsOfSize.{w} G where - preservesColimitsOfShape := preservesColimitOfPreservesCoequalizersAndCoproduct G + [∀ J, PreservesColimitsOfShape (Discrete.{w} J) G] : PreservesColimitsOfSize.{w, w} G where + preservesColimitsOfShape := preservesColimit_of_preservesCoequalizers_and_coproduct G theorem hasFiniteColimits_of_hasInitial_and_pushouts [HasInitial C] [HasPushouts C] : HasFiniteColimits C := @@ -451,16 +450,16 @@ theorem hasFiniteColimits_of_hasInitial_and_pushouts [HasInitial C] [HasPushouts (hasBinaryCoproducts_of_hasInitial_and_pushouts C) inferInstance) /-- If G preserves initial objects and pushouts, it preserves all finite colimits. -/ -noncomputable def preservesFiniteColimitsOfPreservesInitialAndPushouts [HasInitial C] +lemma preservesFiniteColimits_of_preservesInitial_and_pushouts [HasInitial C] [HasPushouts C] (G : C ⥤ D) [PreservesColimitsOfShape (Discrete.{0} PEmpty) G] [PreservesColimitsOfShape WalkingSpan G] : PreservesFiniteColimits G := by haveI : HasFiniteColimits C := hasFiniteColimits_of_hasInitial_and_pushouts haveI : PreservesColimitsOfShape (Discrete WalkingPair) G := - preservesBinaryCoproductsOfPreservesInitialAndPushouts G + preservesBinaryCoproducts_of_preservesInitial_and_pushouts G haveI : PreservesColimitsOfShape (WalkingParallelPair) G := - (preservesCoequalizersOfPreservesPushoutsAndBinaryCoproducts G) + (preservesCoequalizers_of_preservesPushouts_and_binaryCoproducts G) refine - @preservesFiniteColimitsOfPreservesCoequalizersAndFiniteCoproducts _ _ _ _ _ _ G _ ?_ + @preservesFiniteColimits_of_preservesCoequalizers_and_finiteCoproducts _ _ _ _ _ _ G _ ?_ apply PreservesFiniteCoproducts.mk apply preservesFiniteCoproductsOfPreservesBinaryAndInitial G diff --git a/Mathlib/CategoryTheory/Limits/Creates.lean b/Mathlib/CategoryTheory/Limits/Creates.lean index 3b22aa0a3b6f5..148eafbe8efa1 100644 --- a/Mathlib/CategoryTheory/Limits/Creates.lean +++ b/Mathlib/CategoryTheory/Limits/Creates.lean @@ -141,7 +141,7 @@ lemma liftedLimitMapsToOriginal_inv_map_π /-- The lifted cone is a limit. -/ def liftedLimitIsLimit {K : J ⥤ C} {F : C ⥤ D} [CreatesLimit K F] {c : Cone (K ⋙ F)} (t : IsLimit c) : IsLimit (liftLimit t) := - ReflectsLimit.reflects (IsLimit.ofIsoLimit t (liftedLimitMapsToOriginal t).symm) + isLimitOfReflects _ (IsLimit.ofIsoLimit t (liftedLimitMapsToOriginal t).symm) /-- If `F` creates the limit of `K` and `K ⋙ F` has a limit, then `K` has a limit. -/ theorem hasLimit_of_created (K : J ⥤ C) (F : C ⥤ D) [HasLimit (K ⋙ F)] [CreatesLimit K F] : @@ -176,7 +176,7 @@ def liftedColimitMapsToOriginal {K : J ⥤ C} {F : C ⥤ D} [CreatesColimit K F] /-- The lifted cocone is a colimit. -/ def liftedColimitIsColimit {K : J ⥤ C} {F : C ⥤ D} [CreatesColimit K F] {c : Cocone (K ⋙ F)} (t : IsColimit c) : IsColimit (liftColimit t) := - ReflectsColimit.reflects (IsColimit.ofIsoColimit t (liftedColimitMapsToOriginal t).symm) + isColimitOfReflects _ (IsColimit.ofIsoColimit t (liftedColimitMapsToOriginal t).symm) /-- If `F` creates the limit of `K` and `K ⋙ F` has a limit, then `K` has a limit. -/ theorem hasColimit_of_created (K : J ⥤ C) (F : C ⥤ D) [HasColimit (K ⋙ F)] [CreatesColimit K F] : @@ -239,7 +239,7 @@ def createsLimitOfReflectsIso {K : J ⥤ C} {F : C ⥤ D} [F.ReflectsIsomorphism (h : ∀ c t, LiftsToLimit K F c t) : CreatesLimit K F where lifts c t := (h c t).toLiftableCone toReflectsLimit := - { reflects := fun {d} hd => by + { reflects := fun {d} hd => ⟨by let d' : Cone K := (h (F.mapCone d) hd).toLiftableCone.liftedCone let i : F.mapCone d' ≅ F.mapCone d := (h (F.mapCone d) hd).toLiftableCone.validLift @@ -251,7 +251,7 @@ def createsLimitOfReflectsIso {K : J ⥤ C} {F : C ⥤ D} [F.ReflectsIsomorphism rw [this] infer_instance haveI : IsIso f := isIso_of_reflects_iso f (Cones.functoriality K F) - exact IsLimit.ofIsoLimit hd' (asIso f).symm } + exact IsLimit.ofIsoLimit hd' (asIso f).symm⟩ } /-- If `F` reflects isomorphisms and we can lift a single limit cone to a limit cone, then `F` creates limits. Note that unlike `createsLimitOfReflectsIso`, to apply this result it is @@ -322,27 +322,44 @@ def createsLimitOfFullyFaithfulOfIso {K : J ⥤ C} {F : C ⥤ D} [F.Full] [F.Fai /-- A fully faithful functor that preserves a limit that exists also creates the limit. -/ def createsLimitOfFullyFaithfulOfPreserves {K : J ⥤ C} {F : C ⥤ D} [F.Full] [F.Faithful] [HasLimit K] [PreservesLimit K F] : CreatesLimit K F := - createsLimitOfFullyFaithfulOfLift' (PreservesLimit.preserves <| limit.isLimit K) _ (Iso.refl _) + createsLimitOfFullyFaithfulOfLift' (isLimitOfPreserves _ (limit.isLimit K)) _ (Iso.refl _) -- see Note [lower instance priority] /-- `F` preserves the limit of `K` if it creates the limit and `K ⋙ F` has the limit. -/ -instance (priority := 100) preservesLimitOfCreatesLimitAndHasLimit (K : J ⥤ C) (F : C ⥤ D) +instance (priority := 100) preservesLimit_of_createsLimit_and_hasLimit (K : J ⥤ C) (F : C ⥤ D) [CreatesLimit K F] [HasLimit (K ⋙ F)] : PreservesLimit K F where - preserves t := IsLimit.ofIsoLimit (limit.isLimit _) + preserves t := ⟨IsLimit.ofIsoLimit (limit.isLimit _) ((liftedLimitMapsToOriginal (limit.isLimit _)).symm ≪≫ - (Cones.functoriality K F).mapIso ((liftedLimitIsLimit (limit.isLimit _)).uniqueUpToIso t)) + (Cones.functoriality K F).mapIso ((liftedLimitIsLimit (limit.isLimit _)).uniqueUpToIso t))⟩ + +@[deprecated (since := "2024-11-19")] +lemma preservesLimitOfCreatesLimitAndHasLimit (K : J ⥤ C) (F : C ⥤ D) + [CreatesLimit K F] [HasLimit (K ⋙ F)] : PreservesLimit K F := + preservesLimit_of_createsLimit_and_hasLimit _ _ -- see Note [lower instance priority] /-- `F` preserves the limit of shape `J` if it creates these limits and `D` has them. -/ -instance (priority := 100) preservesLimitOfShapeOfCreatesLimitsOfShapeAndHasLimitsOfShape +instance (priority := 100) preservesLimitOfShape_of_createsLimitsOfShape_and_hasLimitsOfShape (F : C ⥤ D) [CreatesLimitsOfShape J F] [HasLimitsOfShape J D] : PreservesLimitsOfShape J F where +@[deprecated (since := "2024-11-19")] +lemma preservesLimitOfShapeOfCreatesLimitsOfShapeAndHasLimitsOfShape + (F : C ⥤ D) [CreatesLimitsOfShape J F] [HasLimitsOfShape J D] : + PreservesLimitsOfShape J F := + preservesLimitOfShape_of_createsLimitsOfShape_and_hasLimitsOfShape _ + -- see Note [lower instance priority] /-- `F` preserves limits if it creates limits and `D` has limits. -/ -instance (priority := 100) preservesLimitsOfCreatesLimitsAndHasLimits (F : C ⥤ D) +instance (priority := 100) preservesLimits_of_createsLimits_and_hasLimits (F : C ⥤ D) [CreatesLimitsOfSize.{w, w'} F] [HasLimitsOfSize.{w, w'} D] : PreservesLimitsOfSize.{w, w'} F where +@[deprecated (since := "2024-11-19")] +lemma preservesLimitsOfCreatesLimitsAndHasLimits (F : C ⥤ D) + [CreatesLimitsOfSize.{w, w'} F] [HasLimitsOfSize.{w, w'} D] : + PreservesLimitsOfSize.{w, w'} F := + preservesLimits_of_createsLimits_and_hasLimits _ + /-- If `F` reflects isomorphisms and we can lift any colimit cocone to a colimit cocone, then `F` creates colimits. In particular here we don't need to assume that F reflects colimits. @@ -351,8 +368,7 @@ def createsColimitOfReflectsIso {K : J ⥤ C} {F : C ⥤ D} [F.ReflectsIsomorphi (h : ∀ c t, LiftsToColimit K F c t) : CreatesColimit K F where lifts c t := (h c t).toLiftableCocone toReflectsColimit := - { - reflects := fun {d} hd => by + { reflects := fun {d} hd => ⟨by let d' : Cocone K := (h (F.mapCocone d) hd).toLiftableCocone.liftedCocone let i : F.mapCocone d' ≅ F.mapCocone d := (h (F.mapCocone d) hd).toLiftableCocone.validLift @@ -364,7 +380,7 @@ def createsColimitOfReflectsIso {K : J ⥤ C} {F : C ⥤ D} [F.ReflectsIsomorphi rw [this] infer_instance haveI := isIso_of_reflects_iso f (Cocones.functoriality K F) - exact IsColimit.ofIsoColimit hd' (asIso f) } + exact IsColimit.ofIsoColimit hd' (asIso f)⟩ } /-- If `F` reflects isomorphisms and we can lift a single colimit cocone to a colimit cocone, then `F` creates limits. Note that unlike `createsColimitOfReflectsIso`, to apply this result it is @@ -407,7 +423,7 @@ def createsColimitOfFullyFaithfulOfLift {K : J ⥤ C} {F : C ⥤ D} [F.Full] [F. /-- A fully faithful functor that preserves a colimit that exists also creates the colimit. -/ def createsColimitOfFullyFaithfulOfPreserves {K : J ⥤ C} {F : C ⥤ D} [F.Full] [F.Faithful] [HasColimit K] [PreservesColimit K F] : CreatesColimit K F := - createsColimitOfFullyFaithfulOfLift' (PreservesColimit.preserves <| colimit.isColimit K) _ + createsColimitOfFullyFaithfulOfLift' (isColimitOfPreserves _ (colimit.isColimit K)) _ (Iso.refl _) -- Notice however that even if the isomorphism is `Iso.refl _`, @@ -442,30 +458,47 @@ def createsColimitOfFullyFaithfulOfIso {K : J ⥤ C} {F : C ⥤ D} [F.Full] [F.F -- see Note [lower instance priority] /-- `F` preserves the colimit of `K` if it creates the colimit and `K ⋙ F` has the colimit. -/ -instance (priority := 100) preservesColimitOfCreatesColimitAndHasColimit (K : J ⥤ C) (F : C ⥤ D) +instance (priority := 100) preservesColimit_of_createsColimit_and_hasColimit (K : J ⥤ C) (F : C ⥤ D) [CreatesColimit K F] [HasColimit (K ⋙ F)] : PreservesColimit K F where preserves t := - IsColimit.ofIsoColimit (colimit.isColimit _) + ⟨IsColimit.ofIsoColimit (colimit.isColimit _) ((liftedColimitMapsToOriginal (colimit.isColimit _)).symm ≪≫ (Cocones.functoriality K F).mapIso - ((liftedColimitIsColimit (colimit.isColimit _)).uniqueUpToIso t)) + ((liftedColimitIsColimit (colimit.isColimit _)).uniqueUpToIso t))⟩ + +@[deprecated (since := "2024-11-19")] +lemma preservesColimitOfCreatesColimitAndHasColimit (K : J ⥤ C) (F : C ⥤ D) + [CreatesColimit K F] [HasColimit (K ⋙ F)] : PreservesColimit K F := + preservesColimit_of_createsColimit_and_hasColimit _ _ -- see Note [lower instance priority] /-- `F` preserves the colimit of shape `J` if it creates these colimits and `D` has them. -/ -instance (priority := 100) preservesColimitOfShapeOfCreatesColimitsOfShapeAndHasColimitsOfShape +instance (priority := 100) preservesColimitOfShape_of_createsColimitsOfShape_and_hasColimitsOfShape (F : C ⥤ D) [CreatesColimitsOfShape J F] [HasColimitsOfShape J D] : PreservesColimitsOfShape J F where +@[deprecated (since := "2024-11-19")] +lemma preservesColimitOfShapeOfCreatesColimitsOfShapeAndHasColimitsOfShape + (F : C ⥤ D) [CreatesColimitsOfShape J F] [HasColimitsOfShape J D] : + PreservesColimitsOfShape J F := + preservesColimitOfShape_of_createsColimitsOfShape_and_hasColimitsOfShape _ + -- see Note [lower instance priority] /-- `F` preserves limits if it creates limits and `D` has limits. -/ -instance (priority := 100) preservesColimitsOfCreatesColimitsAndHasColimits (F : C ⥤ D) +instance (priority := 100) preservesColimits_of_createsColimits_and_hasColimits (F : C ⥤ D) [CreatesColimitsOfSize.{w, w'} F] [HasColimitsOfSize.{w, w'} D] : PreservesColimitsOfSize.{w, w'} F where +@[deprecated (since := "2024-11-19")] +lemma preservesColimitsOfCreatesColimitsAndHasColimits (F : C ⥤ D) + [CreatesColimitsOfSize.{w, w'} F] [HasColimitsOfSize.{w, w'} D] : + PreservesColimitsOfSize.{w, w'} F := + preservesColimits_of_createsColimits_and_hasColimits _ + /-- Transfer creation of limits along a natural isomorphism in the diagram. -/ def createsLimitOfIsoDiagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) [CreatesLimit K₁ F] : CreatesLimit K₂ F := - { reflectsLimitOfIsoDiagram F h with + { reflectsLimit_of_iso_diagram F h with lifts := fun c t => let t' := (IsLimit.postcomposeInvEquiv (isoWhiskerRight h F : _) c).symm t { liftedCone := (Cones.postcompose h.hom).obj (liftLimit t') @@ -485,7 +518,7 @@ def createsLimitOfNatIso {F G : C ⥤ D} (h : F ≅ G) [CreatesLimit K F] : Crea refine (IsLimit.mapConeEquiv h ?_).uniqueUpToIso t apply IsLimit.ofIsoLimit _ (liftedLimitMapsToOriginal _).symm apply (IsLimit.postcomposeInvEquiv _ _).symm t } - toReflectsLimit := reflectsLimitOfNatIso _ h + toReflectsLimit := reflectsLimit_of_natIso _ h /-- If `F` creates limits of shape `J` and `F ≅ G`, then `G` creates limits of shape `J`. -/ def createsLimitsOfShapeOfNatIso {F G : C ⥤ D} (h : F ≅ G) [CreatesLimitsOfShape J F] : @@ -499,7 +532,7 @@ def createsLimitsOfNatIso {F G : C ⥤ D} (h : F ≅ G) [CreatesLimitsOfSize.{w, /-- Transfer creation of colimits along a natural isomorphism in the diagram. -/ def createsColimitOfIsoDiagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) [CreatesColimit K₁ F] : CreatesColimit K₂ F := - { reflectsColimitOfIsoDiagram F h with + { reflectsColimit_of_iso_diagram F h with lifts := fun c t => let t' := (IsColimit.precomposeHomEquiv (isoWhiskerRight h F : _) c).symm t { liftedCocone := (Cocones.precompose h.inv).obj (liftColimit t') @@ -520,7 +553,7 @@ def createsColimitOfNatIso {F G : C ⥤ D} (h : F ≅ G) [CreatesColimit K F] : refine (IsColimit.mapCoconeEquiv h ?_).uniqueUpToIso t apply IsColimit.ofIsoColimit _ (liftedColimitMapsToOriginal _).symm apply (IsColimit.precomposeHomEquiv _ _).symm t } - toReflectsColimit := reflectsColimitOfNatIso _ h + toReflectsColimit := reflectsColimit_of_natIso _ h /-- If `F` creates colimits of shape `J` and `F ≅ G`, then `G` creates colimits of shape `J`. -/ def createsColimitsOfShapeOfNatIso {F G : C ⥤ D} (h : F ≅ G) [CreatesColimitsOfShape J F] : diff --git a/Mathlib/CategoryTheory/Limits/ExactFunctor.lean b/Mathlib/CategoryTheory/Limits/ExactFunctor.lean index cb283bb686903..811252d88fc96 100644 --- a/Mathlib/CategoryTheory/Limits/ExactFunctor.lean +++ b/Mathlib/CategoryTheory/Limits/ExactFunctor.lean @@ -31,7 +31,7 @@ variable (C) (D) -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- Bundled left-exact functors. -/ def LeftExactFunctor := - FullSubcategory fun F : C ⥤ D => Nonempty (PreservesFiniteLimits F) + FullSubcategory fun F : C ⥤ D => PreservesFiniteLimits F instance : Category (LeftExactFunctor C D) := FullSubcategory.category _ @@ -52,7 +52,7 @@ instance : (LeftExactFunctor.forget C D).Faithful := -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- Bundled right-exact functors. -/ def RightExactFunctor := - FullSubcategory fun F : C ⥤ D => Nonempty (PreservesFiniteColimits F) + FullSubcategory fun F : C ⥤ D => PreservesFiniteColimits F instance : Category (RightExactFunctor C D) := FullSubcategory.category _ @@ -74,7 +74,7 @@ instance : (RightExactFunctor.forget C D).Faithful := /-- Bundled exact functors. -/ def ExactFunctor := FullSubcategory fun F : C ⥤ D => - Nonempty (PreservesFiniteLimits F) ∧ Nonempty (PreservesFiniteColimits F) + PreservesFiniteLimits F ∧ PreservesFiniteColimits F instance : Category (ExactFunctor C D) := FullSubcategory.category _ @@ -162,15 +162,15 @@ theorem ExactFunctor.forget_map {F G : C ⥤ₑ D} (α : F ⟶ G) : (ExactFuncto /-- Turn a left exact functor into an object of the category `LeftExactFunctor C D`. -/ def LeftExactFunctor.of (F : C ⥤ D) [PreservesFiniteLimits F] : C ⥤ₗ D := - ⟨F, ⟨inferInstance⟩⟩ + ⟨F, inferInstance⟩ /-- Turn a right exact functor into an object of the category `RightExactFunctor C D`. -/ def RightExactFunctor.of (F : C ⥤ D) [PreservesFiniteColimits F] : C ⥤ᵣ D := - ⟨F, ⟨inferInstance⟩⟩ + ⟨F, inferInstance⟩ /-- Turn an exact functor into an object of the category `ExactFunctor C D`. -/ def ExactFunctor.of (F : C ⥤ D) [PreservesFiniteLimits F] [PreservesFiniteColimits F] : C ⥤ₑ D := - ⟨F, ⟨⟨inferInstance⟩, ⟨inferInstance⟩⟩⟩ + ⟨F, ⟨inferInstance, inferInstance⟩⟩ @[simp] theorem LeftExactFunctor.of_fst (F : C ⥤ D) [PreservesFiniteLimits F] : @@ -200,16 +200,16 @@ theorem ExactFunctor.forget_obj_of (F : C ⥤ D) [PreservesFiniteLimits F] rfl noncomputable instance (F : C ⥤ₗ D) : PreservesFiniteLimits F.obj := - F.property.some + F.property noncomputable instance (F : C ⥤ᵣ D) : PreservesFiniteColimits F.obj := - F.property.some + F.property noncomputable instance (F : C ⥤ₑ D) : PreservesFiniteLimits F.obj := - F.property.1.some + F.property.1 noncomputable instance (F : C ⥤ₑ D) : PreservesFiniteColimits F.obj := - F.property.2.some + F.property.2 variable {E : Type u₃} [Category.{v₃} E] @@ -221,7 +221,7 @@ variable (C D E) @[simps!] def LeftExactFunctor.whiskeringLeft : (C ⥤ₗ D) ⥤ (D ⥤ₗ E) ⥤ (C ⥤ₗ E) where obj F := FullSubcategory.lift _ (forget _ _ ⋙ (CategoryTheory.whiskeringLeft C D E).obj F.obj) - (fun G => ⟨by dsimp; exact compPreservesFiniteLimits _ _⟩) + (fun G => by dsimp; exact comp_preservesFiniteLimits _ _) map {F G} η := { app := fun H => ((CategoryTheory.whiskeringLeft C D E).map η).app H.obj naturality := fun _ _ f => ((CategoryTheory.whiskeringLeft C D E).map η).naturality f } @@ -236,7 +236,7 @@ def LeftExactFunctor.whiskeringLeft : (C ⥤ₗ D) ⥤ (D ⥤ₗ E) ⥤ (C ⥤ @[simps!] def LeftExactFunctor.whiskeringRight : (D ⥤ₗ E) ⥤ (C ⥤ₗ D) ⥤ (C ⥤ₗ E) where obj F := FullSubcategory.lift _ (forget _ _ ⋙ (CategoryTheory.whiskeringRight C D E).obj F.obj) - (fun G => ⟨by dsimp; exact compPreservesFiniteLimits _ _⟩) + (fun G => by dsimp; exact comp_preservesFiniteLimits _ _) map {F G} η := { app := fun H => ((CategoryTheory.whiskeringRight C D E).map η).app H.obj naturality := fun _ _ f => ((CategoryTheory.whiskeringRight C D E).map η).naturality f } @@ -251,7 +251,7 @@ def LeftExactFunctor.whiskeringRight : (D ⥤ₗ E) ⥤ (C ⥤ₗ D) ⥤ (C ⥤ @[simps!] def RightExactFunctor.whiskeringLeft : (C ⥤ᵣ D) ⥤ (D ⥤ᵣ E) ⥤ (C ⥤ᵣ E) where obj F := FullSubcategory.lift _ (forget _ _ ⋙ (CategoryTheory.whiskeringLeft C D E).obj F.obj) - (fun G => ⟨by dsimp; exact compPreservesFiniteColimits _ _⟩) + (fun G => by dsimp; exact comp_preservesFiniteColimits _ _) map {F G} η := { app := fun H => ((CategoryTheory.whiskeringLeft C D E).map η).app H.obj naturality := fun _ _ f => ((CategoryTheory.whiskeringLeft C D E).map η).naturality f } @@ -266,7 +266,7 @@ def RightExactFunctor.whiskeringLeft : (C ⥤ᵣ D) ⥤ (D ⥤ᵣ E) ⥤ (C ⥤ @[simps!] def RightExactFunctor.whiskeringRight : (D ⥤ᵣ E) ⥤ (C ⥤ᵣ D) ⥤ (C ⥤ᵣ E) where obj F := FullSubcategory.lift _ (forget _ _ ⋙ (CategoryTheory.whiskeringRight C D E).obj F.obj) - (fun G => ⟨by dsimp; exact compPreservesFiniteColimits _ _⟩) + (fun G => by dsimp; exact comp_preservesFiniteColimits _ _) map {F G} η := { app := fun H => ((CategoryTheory.whiskeringRight C D E).map η).app H.obj naturality := fun _ _ f => ((CategoryTheory.whiskeringRight C D E).map η).naturality f } @@ -281,8 +281,8 @@ def RightExactFunctor.whiskeringRight : (D ⥤ᵣ E) ⥤ (C ⥤ᵣ D) ⥤ (C ⥤ @[simps!] def ExactFunctor.whiskeringLeft : (C ⥤ₑ D) ⥤ (D ⥤ₑ E) ⥤ (C ⥤ₑ E) where obj F := FullSubcategory.lift _ (forget _ _ ⋙ (CategoryTheory.whiskeringLeft C D E).obj F.obj) - (fun G => ⟨⟨by dsimp; exact compPreservesFiniteLimits _ _⟩, - ⟨by dsimp; exact compPreservesFiniteColimits _ _⟩⟩) + (fun G => ⟨by dsimp; exact comp_preservesFiniteLimits _ _, + by dsimp; exact comp_preservesFiniteColimits _ _⟩) map {F G} η := { app := fun H => ((CategoryTheory.whiskeringLeft C D E).map η).app H.obj naturality := fun _ _ f => ((CategoryTheory.whiskeringLeft C D E).map η).naturality f } @@ -297,8 +297,8 @@ def ExactFunctor.whiskeringLeft : (C ⥤ₑ D) ⥤ (D ⥤ₑ E) ⥤ (C ⥤ₑ E) @[simps!] def ExactFunctor.whiskeringRight : (D ⥤ₑ E) ⥤ (C ⥤ₑ D) ⥤ (C ⥤ₑ E) where obj F := FullSubcategory.lift _ (forget _ _ ⋙ (CategoryTheory.whiskeringRight C D E).obj F.obj) - (fun G => ⟨⟨by dsimp; exact compPreservesFiniteLimits _ _⟩, - ⟨by dsimp; exact compPreservesFiniteColimits _ _⟩⟩) + (fun G => ⟨by dsimp; exact comp_preservesFiniteLimits _ _, + by dsimp; exact comp_preservesFiniteColimits _ _⟩) map {F G} η := { app := fun H => ((CategoryTheory.whiskeringRight C D E).map η).app H.obj naturality := fun _ _ f => ((CategoryTheory.whiskeringRight C D E).map η).naturality f } diff --git a/Mathlib/CategoryTheory/Limits/FilteredColimitCommutesFiniteLimit.lean b/Mathlib/CategoryTheory/Limits/FilteredColimitCommutesFiniteLimit.lean index 1702595d0a0b5..494fe94ffdb41 100644 --- a/Mathlib/CategoryTheory/Limits/FilteredColimitCommutesFiniteLimit.lean +++ b/Mathlib/CategoryTheory/Limits/FilteredColimitCommutesFiniteLimit.lean @@ -338,11 +338,11 @@ instance colimitLimitToLimitColimitCone_iso (F : J ⥤ K ⥤ Type v) : infer_instance apply Cones.cone_iso_of_hom_iso -noncomputable instance filteredColimPreservesFiniteLimitsOfTypes : +noncomputable instance filtered_colim_preservesFiniteLimits_of_types : PreservesFiniteLimits (colim : (K ⥤ Type v) ⥤ _) := by - apply preservesFiniteLimitsOfPreservesFiniteLimitsOfSize.{v₂} + apply preservesFiniteLimits_of_preservesFiniteLimitsOfSize.{v₂} intro J _ _ - refine ⟨fun {F} => ⟨fun {c} hc => IsLimit.ofIsoLimit (limit.isLimit _) ?_⟩⟩ + refine ⟨fun {F} => ⟨fun {c} hc => ⟨IsLimit.ofIsoLimit (limit.isLimit _) ?_⟩⟩⟩ symm trans colim.mapCone (limit.cone F) · exact Functor.mapIso _ (hc.uniqueUpToIso (limit.isLimit F)) @@ -356,20 +356,20 @@ variable [HasLimitsOfShape J C] [HasColimitsOfShape K C] variable [ReflectsLimitsOfShape J (forget C)] [PreservesColimitsOfShape K (forget C)] variable [PreservesLimitsOfShape J (forget C)] -noncomputable instance filteredColimPreservesFiniteLimits : +noncomputable instance filtered_colim_preservesFiniteLimits : PreservesLimitsOfShape J (colim : (K ⥤ C) ⥤ _) := haveI : PreservesLimitsOfShape J ((colim : (K ⥤ C) ⥤ _) ⋙ forget C) := - preservesLimitsOfShapeOfNatIso (preservesColimitNatIso _).symm - preservesLimitsOfShapeOfReflectsOfPreserves _ (forget C) + preservesLimitsOfShape_of_natIso (preservesColimitNatIso _).symm + preservesLimitsOfShape_of_reflects_of_preserves _ (forget C) end -attribute [local instance] reflectsLimitsOfShapeOfReflectsIsomorphisms +attribute [local instance] reflectsLimitsOfShape_of_reflectsIsomorphisms noncomputable instance [PreservesFiniteLimits (forget C)] [PreservesColimitsOfShape K (forget C)] [HasFiniteLimits C] [HasColimitsOfShape K C] [(forget C).ReflectsIsomorphisms] : PreservesFiniteLimits (colim : (K ⥤ C) ⥤ _) := by - apply preservesFiniteLimitsOfPreservesFiniteLimitsOfSize.{v} + apply preservesFiniteLimits_of_preservesFiniteLimitsOfSize.{v} intro J _ _ infer_instance diff --git a/Mathlib/CategoryTheory/Limits/FintypeCat.lean b/Mathlib/CategoryTheory/Limits/FintypeCat.lean index d856af586c75e..a326c743aabf4 100644 --- a/Mathlib/CategoryTheory/Limits/FintypeCat.lean +++ b/Mathlib/CategoryTheory/Limits/FintypeCat.lean @@ -51,14 +51,14 @@ instance {J : Type} [SmallCategory J] [FinCategory J] : HasLimitsOfShape J Finty instance hasFiniteLimits : HasFiniteLimits FintypeCat.{u} where out _ := inferInstance -noncomputable instance inclusionPreservesFiniteLimits : +noncomputable instance inclusion_preservesFiniteLimits : PreservesFiniteLimits FintypeCat.incl.{u} where preservesFiniteLimits _ := - preservesLimitOfShapeOfCreatesLimitsOfShapeAndHasLimitsOfShape FintypeCat.incl + preservesLimitOfShape_of_createsLimitsOfShape_and_hasLimitsOfShape FintypeCat.incl /- Help typeclass inference to infer preservation of finite limits for the forgtful functor. -/ noncomputable instance : PreservesFiniteLimits (forget FintypeCat) := - FintypeCat.inclusionPreservesFiniteLimits + FintypeCat.inclusion_preservesFiniteLimits /-- The categorical product of a finite family in `FintypeCat` is equivalent to the product as types. -/ @@ -112,14 +112,14 @@ instance {J : Type} [SmallCategory J] [FinCategory J] : HasColimitsOfShape J Fin instance hasFiniteColimits : HasFiniteColimits FintypeCat.{u} where out _ := inferInstance -noncomputable instance inclusionPreservesFiniteColimits : +noncomputable instance inclusion_preservesFiniteColimits : PreservesFiniteColimits FintypeCat.incl.{u} where preservesFiniteColimits _ := - preservesColimitOfShapeOfCreatesColimitsOfShapeAndHasColimitsOfShape FintypeCat.incl + preservesColimitOfShape_of_createsColimitsOfShape_and_hasColimitsOfShape FintypeCat.incl /- Help typeclass inference to infer preservation of finite colimits for the forgtful functor. -/ noncomputable instance : PreservesFiniteColimits (forget FintypeCat) := - FintypeCat.inclusionPreservesFiniteColimits + FintypeCat.inclusion_preservesFiniteColimits lemma jointly_surjective {J : Type*} [Category J] [FinCategory J] (F : J ⥤ FintypeCat.{u}) (t : Cocone F) (h : IsColimit t) (x : t.pt) : diff --git a/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean b/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean index 86b8cb71e4d54..357916b1fe01e 100644 --- a/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean +++ b/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean @@ -12,12 +12,12 @@ import Mathlib.CategoryTheory.Limits.Preserves.Limits We show that if `D` has limits, then the functor category `C ⥤ D` also has limits (`CategoryTheory.Limits.functorCategoryHasLimits`), and the evaluation functors preserve limits -(`CategoryTheory.Limits.evaluationPreservesLimits`) +(`CategoryTheory.Limits.evaluation_preservesLimits`) (and similarly for colimits). We also show that `F : D ⥤ K ⥤ C` preserves (co)limits if it does so for each `k : K` -(`CategoryTheory.Limits.preservesLimitsOfEvaluation` and -`CategoryTheory.Limits.preservesColimitsOfEvaluation`). +(`CategoryTheory.Limits.preservesLimits_of_evaluation` and +`CategoryTheory.Limits.preservesColimits_of_evaluation`). -/ @@ -178,14 +178,14 @@ instance hasLimitCompEvalution (F : J ⥤ K ⥤ C) (k : K) [HasLimit (F.flip.obj HasLimit (F ⋙ (evaluation _ _).obj k) := hasLimitOfIso (F := F.flip.obj k) (Iso.refl _) -instance evaluationPreservesLimit (F : J ⥤ K ⥤ C) [∀ k, HasLimit (F.flip.obj k)] (k : K) : +instance evaluation_preservesLimit (F : J ⥤ K ⥤ C) [∀ k, HasLimit (F.flip.obj k)] (k : K) : PreservesLimit F ((evaluation K C).obj k) := -- Porting note: added a let because X was not inferred let X : (k : K) → LimitCone (F.flip.obj k) := fun k => getLimitCone (F.flip.obj k) - preservesLimitOfPreservesLimitCone (combinedIsLimit _ X) <| + preservesLimit_of_preserves_limit_cone (combinedIsLimit _ X) <| IsLimit.ofIsoLimit (limit.isLimit _) (evaluateCombinedCones F X k).symm -instance evaluationPreservesLimitsOfShape [HasLimitsOfShape J C] (k : K) : +instance evaluation_preservesLimitsOfShape [HasLimitsOfShape J C] (k : K) : PreservesLimitsOfShape J ((evaluation K C).obj k) where preservesLimit := inferInstance @@ -263,14 +263,14 @@ instance hasColimitCompEvaluation (F : J ⥤ K ⥤ C) (k : K) [HasColimit (F.fli HasColimit (F ⋙ (evaluation _ _).obj k) := hasColimitOfIso (F := F.flip.obj k) (Iso.refl _) -instance evaluationPreservesColimit (F : J ⥤ K ⥤ C) [∀ k, HasColimit (F.flip.obj k)] (k : K) : +instance evaluation_preservesColimit (F : J ⥤ K ⥤ C) [∀ k, HasColimit (F.flip.obj k)] (k : K) : PreservesColimit F ((evaluation K C).obj k) := -- Porting note: added a let because X was not inferred let X : (k : K) → ColimitCocone (F.flip.obj k) := fun k => getColimitCocone (F.flip.obj k) - preservesColimitOfPreservesColimitCocone (combinedIsColimit _ X) <| + preservesColimit_of_preserves_colimit_cocone (combinedIsColimit _ X) <| IsColimit.ofIsoColimit (colimit.isColimit _) (evaluateCombinedCocones F X k).symm -instance evaluationPreservesColimitsOfShape [HasColimitsOfShape J C] (k : K) : +instance evaluation_preservesColimitsOfShape [HasColimitsOfShape J C] (k : K) : PreservesColimitsOfShape J ((evaluation K C).obj k) where preservesColimit := inferInstance @@ -353,65 +353,88 @@ instance evaluationPreservesLimits [HasLimits C] (k : K) : preservesLimitsOfShape {_} _𝒥 := inferInstance /-- `F : D ⥤ K ⥤ C` preserves the limit of some `G : J ⥤ D` if it does for each `k : K`. -/ -def preservesLimitOfEvaluation (F : D ⥤ K ⥤ C) (G : J ⥤ D) +lemma preservesLimit_of_evaluation (F : D ⥤ K ⥤ C) (G : J ⥤ D) (H : ∀ k : K, PreservesLimit G (F ⋙ (evaluation K C).obj k : D ⥤ C)) : PreservesLimit G F := - ⟨fun {c} hc => by + ⟨fun {c} hc => ⟨by apply evaluationJointlyReflectsLimits intro X haveI := H X change IsLimit ((F ⋙ (evaluation K C).obj X).mapCone c) - exact PreservesLimit.preserves hc⟩ + exact isLimitOfPreserves _ hc⟩⟩ + +@[deprecated (since := "2024-11-19")] +lemma preservesLimitOfEvaluation (F : D ⥤ K ⥤ C) (G : J ⥤ D) + (H : ∀ k : K, PreservesLimit G (F ⋙ (evaluation K C).obj k : D ⥤ C)) : + PreservesLimit G F := + preservesLimit_of_evaluation _ _ H /-- `F : D ⥤ K ⥤ C` preserves limits of shape `J` if it does for each `k : K`. -/ -def preservesLimitsOfShapeOfEvaluation (F : D ⥤ K ⥤ C) (J : Type*) [Category J] +lemma preservesLimitsOfShape_of_evaluation (F : D ⥤ K ⥤ C) (J : Type*) [Category J] (_ : ∀ k : K, PreservesLimitsOfShape J (F ⋙ (evaluation K C).obj k)) : PreservesLimitsOfShape J F := - ⟨fun {G} => preservesLimitOfEvaluation F G fun _ => PreservesLimitsOfShape.preservesLimit⟩ + ⟨fun {G} => preservesLimit_of_evaluation F G fun _ => PreservesLimitsOfShape.preservesLimit⟩ + +@[deprecated (since := "2024-11-19")] +lemma preservesLimitsOfShapeOfEvaluation (F : D ⥤ K ⥤ C) (J : Type*) [Category J] + (H : ∀ k : K, PreservesLimitsOfShape J (F ⋙ (evaluation K C).obj k)) : + PreservesLimitsOfShape J F := + preservesLimitsOfShape_of_evaluation _ _ H /-- `F : D ⥤ K ⥤ C` preserves all limits if it does for each `k : K`. -/ -def preservesLimitsOfEvaluation (F : D ⥤ K ⥤ C) +lemma preservesLimits_of_evaluation (F : D ⥤ K ⥤ C) (_ : ∀ k : K, PreservesLimitsOfSize.{w', w} (F ⋙ (evaluation K C).obj k)) : PreservesLimitsOfSize.{w', w} F := ⟨fun {L} _ => - preservesLimitsOfShapeOfEvaluation F L fun _ => PreservesLimitsOfSize.preservesLimitsOfShape⟩ + preservesLimitsOfShape_of_evaluation F L fun _ => PreservesLimitsOfSize.preservesLimitsOfShape⟩ + +@[deprecated (since := "2024-11-19")] +lemma preservesLimitsOfEvaluation (F : D ⥤ K ⥤ C) + (H : ∀ k : K, PreservesLimitsOfSize.{w', w} (F ⋙ (evaluation K C).obj k)) : + PreservesLimitsOfSize.{w', w} F := + preservesLimits_of_evaluation _ H /-- The constant functor `C ⥤ (D ⥤ C)` preserves limits. -/ -instance preservesLimitsConst : PreservesLimitsOfSize.{w', w} (const D : C ⥤ _) := - preservesLimitsOfEvaluation _ fun _ => - preservesLimitsOfNatIso <| Iso.symm <| constCompEvaluationObj _ _ +instance preservesLimits_const : PreservesLimitsOfSize.{w', w} (const D : C ⥤ _) := + preservesLimits_of_evaluation _ fun _ => + preservesLimits_of_natIso <| Iso.symm <| constCompEvaluationObj _ _ -instance evaluationPreservesColimits [HasColimits C] (k : K) : +instance evaluation_preservesColimits [HasColimits C] (k : K) : PreservesColimits ((evaluation K C).obj k) where preservesColimitsOfShape := inferInstance /-- `F : D ⥤ K ⥤ C` preserves the colimit of some `G : J ⥤ D` if it does for each `k : K`. -/ -def preservesColimitOfEvaluation (F : D ⥤ K ⥤ C) (G : J ⥤ D) +lemma preservesColimit_of_evaluation (F : D ⥤ K ⥤ C) (G : J ⥤ D) (H : ∀ k, PreservesColimit G (F ⋙ (evaluation K C).obj k)) : PreservesColimit G F := - ⟨fun {c} hc => by + ⟨fun {c} hc => ⟨by apply evaluationJointlyReflectsColimits intro X haveI := H X change IsColimit ((F ⋙ (evaluation K C).obj X).mapCocone c) - exact PreservesColimit.preserves hc⟩ + exact isColimitOfPreserves _ hc⟩⟩ + +@[deprecated (since := "2024-11-19")] +lemma preservesColimitOfEvaluation (F : D ⥤ K ⥤ C) (G : J ⥤ D) + (H : ∀ k, PreservesColimit G (F ⋙ (evaluation K C).obj k)) : PreservesColimit G F := + preservesColimit_of_evaluation _ _ H /-- `F : D ⥤ K ⥤ C` preserves all colimits of shape `J` if it does for each `k : K`. -/ -def preservesColimitsOfShapeOfEvaluation (F : D ⥤ K ⥤ C) (J : Type*) [Category J] +lemma preservesColimitsOfShape_of_evaluation (F : D ⥤ K ⥤ C) (J : Type*) [Category J] (_ : ∀ k : K, PreservesColimitsOfShape J (F ⋙ (evaluation K C).obj k)) : PreservesColimitsOfShape J F := - ⟨fun {G} => preservesColimitOfEvaluation F G fun _ => PreservesColimitsOfShape.preservesColimit⟩ + ⟨fun {G} => preservesColimit_of_evaluation F G fun _ => PreservesColimitsOfShape.preservesColimit⟩ /-- `F : D ⥤ K ⥤ C` preserves all colimits if it does for each `k : K`. -/ -def preservesColimitsOfEvaluation (F : D ⥤ K ⥤ C) +lemma preservesColimits_of_evaluation (F : D ⥤ K ⥤ C) (_ : ∀ k : K, PreservesColimitsOfSize.{w', w} (F ⋙ (evaluation K C).obj k)) : PreservesColimitsOfSize.{w', w} F := ⟨fun {L} _ => - preservesColimitsOfShapeOfEvaluation F L fun _ => + preservesColimitsOfShape_of_evaluation F L fun _ => PreservesColimitsOfSize.preservesColimitsOfShape⟩ /-- The constant functor `C ⥤ (D ⥤ C)` preserves colimits. -/ -instance preservesColimitsConst : PreservesColimitsOfSize.{w', w} (const D : C ⥤ _) := - preservesColimitsOfEvaluation _ fun _ => - preservesColimitsOfNatIso <| Iso.symm <| constCompEvaluationObj _ _ +instance preservesColimits_const : PreservesColimitsOfSize.{w', w} (const D : C ⥤ _) := + preservesColimits_of_evaluation _ fun _ => + preservesColimits_of_natIso <| Iso.symm <| constCompEvaluationObj _ _ open CategoryTheory.prod diff --git a/Mathlib/CategoryTheory/Limits/Over.lean b/Mathlib/CategoryTheory/Limits/Over.lean index 34e23b9622e0b..e6a0fb7ce116d 100644 --- a/Mathlib/CategoryTheory/Limits/Over.lean +++ b/Mathlib/CategoryTheory/Limits/Over.lean @@ -63,9 +63,9 @@ instance createsColimitsOfSizeMapCompForget {Y : C} (f : X ⟶ Y) : CreatesColimitsOfSize.{w, w'} (map f ⋙ forget Y) := show CreatesColimitsOfSize.{w, w'} (forget X) from inferInstance -instance preservesColimitsOfSizeMap [HasColimitsOfSize.{w, w'} C] {Y : C} (f : X ⟶ Y) : +instance preservesColimitsOfSize_map [HasColimitsOfSize.{w, w'} C] {Y : C} (f : X ⟶ Y) : PreservesColimitsOfSize.{w, w'} (map f) := - preservesColimitsOfReflectsOfPreserves (map f) (forget Y) + preservesColimits_of_reflects_of_preserves (map f) (forget Y) /-- If `c` is a colimit cocone, then so is the cocone `c.toOver` with cocone point `𝟙 c.pt`. -/ def isColimitToOver {F : J ⥤ C} {c : Cocone F} (hc : IsColimit c) : IsColimit c.toOver := @@ -110,9 +110,9 @@ instance createLimitsOfSizeMapCompForget {Y : C} (f : X ⟶ Y) : CreatesLimitsOfSize.{w, w'} (map f ⋙ forget X) := show CreatesLimitsOfSize.{w, w'} (forget Y) from inferInstance -instance preservesLimitsOfSizeMap [HasLimitsOfSize.{w, w'} C] {Y : C} (f : X ⟶ Y) : +instance preservesLimitsOfSize_map [HasLimitsOfSize.{w, w'} C] {Y : C} (f : X ⟶ Y) : PreservesLimitsOfSize.{w, w'} (map f) := - preservesLimitsOfReflectsOfPreserves (map f) (forget X) + preservesLimits_of_reflects_of_preserves (map f) (forget X) /-- If `c` is a limit cone, then so is the cone `c.toUnder` with cone point `𝟙 c.pt`. -/ def isLimitToUnder {F : J ⥤ C} {c : Cone F} (hc : IsLimit c) : IsLimit c.toUnder := diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Basic.lean b/Mathlib/CategoryTheory/Limits/Preserves/Basic.lean index f5abc3a11449f..c77df5a351b5d 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Basic.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Basic.lean @@ -49,30 +49,30 @@ variable {J : Type w} [Category.{w'} J] {K : J ⥤ C} /-- A functor `F` preserves limits of `K` (written as `PreservesLimit K F`) if `F` maps any limit cone over `K` to a limit cone. -/ -class PreservesLimit (K : J ⥤ C) (F : C ⥤ D) where - preserves : ∀ {c : Cone K}, IsLimit c → IsLimit (F.mapCone c) +class PreservesLimit (K : J ⥤ C) (F : C ⥤ D) : Prop where + preserves {c : Cone K} (hc : IsLimit c) : Nonempty (IsLimit (F.mapCone c)) /-- A functor `F` preserves colimits of `K` (written as `PreservesColimit K F`) if `F` maps any colimit cocone over `K` to a colimit cocone. -/ -class PreservesColimit (K : J ⥤ C) (F : C ⥤ D) where - preserves : ∀ {c : Cocone K}, IsColimit c → IsColimit (F.mapCocone c) +class PreservesColimit (K : J ⥤ C) (F : C ⥤ D) : Prop where + preserves {c : Cocone K} (hc : IsColimit c) : Nonempty (IsColimit (F.mapCocone c)) /-- We say that `F` preserves limits of shape `J` if `F` preserves limits for every diagram `K : J ⥤ C`, i.e., `F` maps limit cones over `K` to limit cones. -/ -class PreservesLimitsOfShape (J : Type w) [Category.{w'} J] (F : C ⥤ D) where +class PreservesLimitsOfShape (J : Type w) [Category.{w'} J] (F : C ⥤ D) : Prop where preservesLimit : ∀ {K : J ⥤ C}, PreservesLimit K F := by infer_instance /-- We say that `F` preserves colimits of shape `J` if `F` preserves colimits for every diagram `K : J ⥤ C`, i.e., `F` maps colimit cocones over `K` to colimit cocones. -/ -class PreservesColimitsOfShape (J : Type w) [Category.{w'} J] (F : C ⥤ D) where +class PreservesColimitsOfShape (J : Type w) [Category.{w'} J] (F : C ⥤ D) : Prop where preservesColimit : ∀ {K : J ⥤ C}, PreservesColimit K F := by infer_instance -- This should be used with explicit universe variables. /-- `PreservesLimitsOfSize.{v u} F` means that `F` sends all limit cones over any diagram `J ⥤ C` to limit cones, where `J : Type u` with `[Category.{v} J]`. -/ @[nolint checkUnivs, pp_with_univ] -class PreservesLimitsOfSize (F : C ⥤ D) where +class PreservesLimitsOfSize (F : C ⥤ D) : Prop where preservesLimitsOfShape : ∀ {J : Type w} [Category.{w'} J], PreservesLimitsOfShape J F := by infer_instance @@ -85,7 +85,7 @@ abbrev PreservesLimits (F : C ⥤ D) := /-- `PreservesColimitsOfSize.{v u} F` means that `F` sends all colimit cocones over any diagram `J ⥤ C` to colimit cocones, where `J : Type u` with `[Category.{v} J]`. -/ @[nolint checkUnivs, pp_with_univ] -class PreservesColimitsOfSize (F : C ⥤ D) where +class PreservesColimitsOfSize (F : C ⥤ D) : Prop where preservesColimitsOfShape : ∀ {J : Type w} [Category.{w'} J], PreservesColimitsOfShape J F := by infer_instance @@ -106,7 +106,7 @@ guide typeclass resolution. -/ def isLimitOfPreserves (F : C ⥤ D) {c : Cone K} (t : IsLimit c) [PreservesLimit K F] : IsLimit (F.mapCone c) := - PreservesLimit.preserves t + (PreservesLimit.preserves t).some /-- A convenience function for `PreservesColimit`, which takes the functor as an explicit argument to @@ -114,33 +114,33 @@ guide typeclass resolution. -/ def isColimitOfPreserves (F : C ⥤ D) {c : Cocone K} (t : IsColimit c) [PreservesColimit K F] : IsColimit (F.mapCocone c) := - PreservesColimit.preserves t + (PreservesColimit.preserves t).some instance preservesLimit_subsingleton (K : J ⥤ C) (F : C ⥤ D) : Subsingleton (PreservesLimit K F) := by - constructor; rintro ⟨a⟩ ⟨b⟩; congr!; subsingleton + constructor; rintro ⟨a⟩ ⟨b⟩; congr! instance preservesColimit_subsingleton (K : J ⥤ C) (F : C ⥤ D) : Subsingleton (PreservesColimit K F) := by - constructor; rintro ⟨a⟩ ⟨b⟩; congr!; subsingleton + constructor; rintro ⟨a⟩ ⟨b⟩; congr! instance preservesLimitsOfShape_subsingleton (J : Type w) [Category.{w'} J] (F : C ⥤ D) : Subsingleton (PreservesLimitsOfShape J F) := by - constructor; rintro ⟨a⟩ ⟨b⟩; congr!; subsingleton + constructor; rintro ⟨a⟩ ⟨b⟩; congr! instance preservesColimitsOfShape_subsingleton (J : Type w) [Category.{w'} J] (F : C ⥤ D) : Subsingleton (PreservesColimitsOfShape J F) := by - constructor; rintro ⟨a⟩ ⟨b⟩; congr!; subsingleton + constructor; rintro ⟨a⟩ ⟨b⟩; congr! -instance preserves_limits_subsingleton (F : C ⥤ D) : +instance preservesLimitsOfSize_subsingleton (F : C ⥤ D) : Subsingleton (PreservesLimitsOfSize.{w', w} F) := by - constructor; rintro ⟨a⟩ ⟨b⟩; congr!; subsingleton + constructor; rintro ⟨a⟩ ⟨b⟩; congr! -instance preserves_colimits_subsingleton (F : C ⥤ D) : +instance preservesColimitsOfSize_subsingleton (F : C ⥤ D) : Subsingleton (PreservesColimitsOfSize.{w', w} F) := by - constructor; rintro ⟨a⟩ ⟨b⟩; congr!; subsingleton + constructor; rintro ⟨a⟩ ⟨b⟩; congr! -instance idPreservesLimits : PreservesLimitsOfSize.{w', w} (𝟭 C) where +instance id_preservesLimitsOfSize : PreservesLimitsOfSize.{w', w} (𝟭 C) where preservesLimitsOfShape {J} 𝒥 := { preservesLimit := fun {K} => @@ -150,7 +150,11 @@ instance idPreservesLimits : PreservesLimitsOfSize.{w', w} (𝟭 C) where cases K; rcases c with ⟨_, _, _⟩; intro s m w; rcases s with ⟨_, _, _⟩; exact h.uniq _ m w⟩⟩ } -instance idPreservesColimits : PreservesColimitsOfSize.{w', w} (𝟭 C) where +@[deprecated "use id_preservesLimitsOfSize" (since := "2024-11-19")] +lemma idPreservesLimits : PreservesLimitsOfSize.{w', w} (𝟭 C) := + id_preservesLimitsOfSize + +instance id_preservesColimitsOfSize : PreservesColimitsOfSize.{w', w} (𝟭 C) where preservesColimitsOfShape {J} 𝒥 := { preservesColimit := fun {K} => @@ -160,12 +164,15 @@ instance idPreservesColimits : PreservesColimitsOfSize.{w', w} (𝟭 C) where cases K; rcases c with ⟨_, _, _⟩; intro s m w; rcases s with ⟨_, _, _⟩; exact h.uniq _ m w⟩⟩ } +@[deprecated "use id_preservesColimitsOfSize" (since := "2024-11-19")] +lemma idPreservesColimits : PreservesColimitsOfSize.{w', w} (𝟭 C) := + id_preservesColimitsOfSize + instance [HasLimit K] {F : C ⥤ D} [PreservesLimit K F] : HasLimit (K ⋙ F) where - exists_limit := ⟨⟨F.mapCone (limit.cone K), PreservesLimit.preserves (limit.isLimit K)⟩⟩ + exists_limit := ⟨_, isLimitOfPreserves F (limit.isLimit K)⟩ instance [HasColimit K] {F : C ⥤ D} [PreservesColimit K F] : HasColimit (K ⋙ F) where - exists_colimit := - ⟨⟨F.mapCocone (colimit.cocone K), PreservesColimit.preserves (colimit.isColimit K)⟩⟩ + exists_colimit := ⟨_, isColimitOfPreserves F (colimit.isColimit K)⟩ section @@ -175,172 +182,296 @@ variable (F : C ⥤ D) (G : D ⥤ E) -- Porting note: made this global by removing local attribute [elab_without_expected_type] PreservesLimit.preserves PreservesColimit.preserves -instance compPreservesLimit [PreservesLimit K F] [PreservesLimit (K ⋙ F) G] : - PreservesLimit K (F ⋙ G) := - ⟨fun h => PreservesLimit.preserves (PreservesLimit.preserves h)⟩ +instance comp_preservesLimit [PreservesLimit K F] [PreservesLimit (K ⋙ F) G] : + PreservesLimit K (F ⋙ G) where + preserves hc := ⟨isLimitOfPreserves G (isLimitOfPreserves F hc)⟩ -instance compPreservesLimitsOfShape [PreservesLimitsOfShape J F] [PreservesLimitsOfShape J G] : +instance comp_preservesLimitsOfShape [PreservesLimitsOfShape J F] [PreservesLimitsOfShape J G] : PreservesLimitsOfShape J (F ⋙ G) where -instance compPreservesLimits [PreservesLimitsOfSize.{w', w} F] [PreservesLimitsOfSize.{w', w} G] : +instance comp_preservesLimits [PreservesLimitsOfSize.{w', w} F] [PreservesLimitsOfSize.{w', w} G] : PreservesLimitsOfSize.{w', w} (F ⋙ G) where -instance compPreservesColimit [PreservesColimit K F] [PreservesColimit (K ⋙ F) G] : - PreservesColimit K (F ⋙ G) := - ⟨fun h => PreservesColimit.preserves (PreservesColimit.preserves h)⟩ +instance comp_preservesColimit [PreservesColimit K F] [PreservesColimit (K ⋙ F) G] : + PreservesColimit K (F ⋙ G) where + preserves hc := ⟨isColimitOfPreserves G (isColimitOfPreserves F hc)⟩ -instance compPreservesColimitsOfShape [PreservesColimitsOfShape J F] +instance comp_preservesColimitsOfShape [PreservesColimitsOfShape J F] [PreservesColimitsOfShape J G] : PreservesColimitsOfShape J (F ⋙ G) where -instance compPreservesColimits [PreservesColimitsOfSize.{w', w} F] +instance comp_preservesColimits [PreservesColimitsOfSize.{w', w} F] [PreservesColimitsOfSize.{w', w} G] : PreservesColimitsOfSize.{w', w} (F ⋙ G) where +@[deprecated "use comp_preservesLimit" (since := "2024-11-19")] +lemma compPreservesLimit [PreservesLimit K F] [PreservesLimit (K ⋙ F) G] : + PreservesLimit K (F ⋙ G) := inferInstance + +@[deprecated "use comp_preservesLimitsOfShape" (since := "2024-11-19")] +lemma compPreservesLimitsOfShape [PreservesLimitsOfShape J F] [PreservesLimitsOfShape J G] : + PreservesLimitsOfShape J (F ⋙ G) := + comp_preservesLimitsOfShape _ _ + +@[deprecated "use comp_preservesLimits" (since := "2024-11-19")] +lemma compPreservesLimits [PreservesLimitsOfSize.{w', w} F] [PreservesLimitsOfSize.{w', w} G] : + PreservesLimitsOfSize.{w', w} (F ⋙ G) := + comp_preservesLimits _ _ + +@[deprecated "use comp_preservesColimit" (since := "2024-11-19")] +lemma compPreservesColimit [PreservesColimit K F] [PreservesColimit (K ⋙ F) G] : + PreservesColimit K (F ⋙ G) := inferInstance + +@[deprecated "use comp_preservesColimitsOfShape" (since := "2024-11-19")] +lemma compPreservesColimitsOfShape [PreservesColimitsOfShape J F] [PreservesColimitsOfShape J G] : + PreservesColimitsOfShape J (F ⋙ G) := + comp_preservesColimitsOfShape _ _ + +@[deprecated "use comp_preservesColimits" (since := "2024-11-19")] +lemma compPreservesColimits [PreservesColimitsOfSize.{w', w} F] + [PreservesColimitsOfSize.{w', w} G] : + PreservesColimitsOfSize.{w', w} (F ⋙ G) := + comp_preservesColimits _ _ + end /-- If F preserves one limit cone for the diagram K, then it preserves any limit cone for K. -/ -def preservesLimitOfPreservesLimitCone {F : C ⥤ D} {t : Cone K} (h : IsLimit t) +lemma preservesLimit_of_preserves_limit_cone {F : C ⥤ D} {t : Cone K} (h : IsLimit t) + (hF : IsLimit (F.mapCone t)) : PreservesLimit K F where + preserves h' := ⟨IsLimit.ofIsoLimit hF (Functor.mapIso _ (IsLimit.uniqueUpToIso h h'))⟩ + +@[deprecated "use preservesLimit_of_preserves_limit_cone" (since := "2024-11-19")] +lemma preservesLimitOfPreservesLimitCone {F : C ⥤ D} {t : Cone K} (h : IsLimit t) (hF : IsLimit (F.mapCone t)) : PreservesLimit K F := - ⟨fun h' => IsLimit.ofIsoLimit hF (Functor.mapIso _ (IsLimit.uniqueUpToIso h h'))⟩ +preservesLimit_of_preserves_limit_cone h hF /-- Transfer preservation of limits along a natural isomorphism in the diagram. -/ -def preservesLimitOfIsoDiagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) [PreservesLimit K₁ F] : - PreservesLimit K₂ F where - preserves {c} t := by +lemma preservesLimit_of_iso_diagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) + [PreservesLimit K₁ F] : PreservesLimit K₂ F where + preserves {c} t := ⟨by apply IsLimit.postcomposeInvEquiv (isoWhiskerRight h F : _) _ _ have := (IsLimit.postcomposeInvEquiv h c).symm t apply IsLimit.ofIsoLimit (isLimitOfPreserves F this) - exact Cones.ext (Iso.refl _) + exact Cones.ext (Iso.refl _)⟩ + +@[deprecated "use preservesLimit_of_iso_diagram" (since := "2024-11-19")] +lemma preservesLimitOfIsoDiagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) + [PreservesLimit K₁ F] : PreservesLimit K₂ F := + preservesLimit_of_iso_diagram F h /-- Transfer preservation of a limit along a natural isomorphism in the functor. -/ -def preservesLimitOfNatIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [PreservesLimit K F] : +lemma preservesLimit_of_natIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [PreservesLimit K F] : PreservesLimit K G where - preserves t := IsLimit.mapConeEquiv h (PreservesLimit.preserves t) + preserves t := ⟨IsLimit.mapConeEquiv h (isLimitOfPreserves F t)⟩ + +@[deprecated "use preservesLimit_of_natIso" (since := "2024-11-19")] +lemma preservesLimitOfNatIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [PreservesLimit K F] : + PreservesLimit K G := + preservesLimit_of_natIso K h /-- Transfer preservation of limits of shape along a natural isomorphism in the functor. -/ -def preservesLimitsOfShapeOfNatIso {F G : C ⥤ D} (h : F ≅ G) [PreservesLimitsOfShape J F] : +lemma preservesLimitsOfShape_of_natIso {F G : C ⥤ D} (h : F ≅ G) [PreservesLimitsOfShape J F] : PreservesLimitsOfShape J G where - preservesLimit {K} := preservesLimitOfNatIso K h + preservesLimit {K} := preservesLimit_of_natIso K h + +@[deprecated "use preservesLimitsOfShape_of_natIso" (since := "2024-11-19")] +lemma preservesLimitsOfShapeOfNatIso {F G : C ⥤ D} (h : F ≅ G) [PreservesLimitsOfShape J F] : + PreservesLimitsOfShape J G := + preservesLimitsOfShape_of_natIso h /-- Transfer preservation of limits along a natural isomorphism in the functor. -/ -def preservesLimitsOfNatIso {F G : C ⥤ D} (h : F ≅ G) [PreservesLimitsOfSize.{w, w'} F] : +lemma preservesLimits_of_natIso {F G : C ⥤ D} (h : F ≅ G) [PreservesLimitsOfSize.{w, w'} F] : PreservesLimitsOfSize.{w, w'} G where - preservesLimitsOfShape {_J} _𝒥₁ := preservesLimitsOfShapeOfNatIso h + preservesLimitsOfShape := preservesLimitsOfShape_of_natIso h + +@[deprecated "use preservesLimits_of_natIso" (since := "2024-11-19")] +lemma preservesLimitsOfNatIso {F G : C ⥤ D} (h : F ≅ G) [PreservesLimitsOfSize.{w, w'} F] : + PreservesLimitsOfSize.{w, w'} G := + preservesLimits_of_natIso h /-- Transfer preservation of limits along an equivalence in the shape. -/ -def preservesLimitsOfShapeOfEquiv {J' : Type w₂} [Category.{w₂'} J'] (e : J ≌ J') (F : C ⥤ D) +lemma preservesLimitsOfShape_of_equiv {J' : Type w₂} [Category.{w₂'} J'] (e : J ≌ J') (F : C ⥤ D) [PreservesLimitsOfShape J F] : PreservesLimitsOfShape J' F where preservesLimit {K} := - { preserves := fun {c} t => by + { preserves := fun {c} t => ⟨by let equ := e.invFunIdAssoc (K ⋙ F) have := (isLimitOfPreserves F (t.whiskerEquivalence e)).whiskerEquivalence e.symm apply ((IsLimit.postcomposeHomEquiv equ _).symm this).ofIsoLimit refine Cones.ext (Iso.refl _) fun j => ?_ dsimp - simp [equ, ← Functor.map_comp] } + simp [equ, ← Functor.map_comp]⟩ } + +@[deprecated "use preservesLimitsOfShape_of_equiv" (since := "2024-11-19")] +lemma preservesLimitsOfShapeOfEquiv {J' : Type w₂} [Category.{w₂'} J'] (e : J ≌ J') (F : C ⥤ D) + [PreservesLimitsOfShape J F] : PreservesLimitsOfShape J' F := + preservesLimitsOfShape_of_equiv e F /-- A functor preserving larger limits also preserves smaller limits. -/ -def preservesLimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] +lemma preservesLimitsOfSize_of_univLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] [PreservesLimitsOfSize.{w', w₂'} F] : PreservesLimitsOfSize.{w, w₂} F where - preservesLimitsOfShape {J} := preservesLimitsOfShapeOfEquiv + preservesLimitsOfShape {J} := preservesLimitsOfShape_of_equiv ((ShrinkHoms.equivalence J).trans <| Shrink.equivalence _).symm F +@[deprecated "use preservesLimitsOfSize_of_univLE" (since := "2024-11-19")] +lemma preservesLimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] + [PreservesLimitsOfSize.{w', w₂'} F] : PreservesLimitsOfSize.{w, w₂} F := + preservesLimitsOfSize_of_univLE.{w', w₂'} F + -- See library note [dsimp, simp]. -/-- `PreservesLimitsOfSizeShrink.{w w'} F` tries to obtain `PreservesLimitsOfSize.{w w'} F` +/-- `PreservesLimitsOfSize_shrink.{w w'} F` tries to obtain `PreservesLimitsOfSize.{w w'} F` from some other `PreservesLimitsOfSize F`. -/ -def preservesLimitsOfSizeShrink (F : C ⥤ D) [PreservesLimitsOfSize.{max w w₂, max w' w₂'} F] : - PreservesLimitsOfSize.{w, w'} F := preservesLimitsOfSizeOfUnivLE.{max w w₂, max w' w₂'} F +lemma preservesLimitsOfSize_shrink (F : C ⥤ D) [PreservesLimitsOfSize.{max w w₂, max w' w₂'} F] : + PreservesLimitsOfSize.{w, w'} F := preservesLimitsOfSize_of_univLE.{max w w₂, max w' w₂'} F + +@[deprecated "use preservesLimitsOfSize_shrink" (since := "2024-11-19")] +lemma PreservesLimitsOfSizeShrink (F : C ⥤ D) [PreservesLimitsOfSize.{max w w₂, max w' w₂'} F] : + PreservesLimitsOfSize.{w, w'} F := + preservesLimitsOfSize_shrink F /-- Preserving limits at any universe level implies preserving limits in universe `0`. -/ -def preservesSmallestLimitsOfPreservesLimits (F : C ⥤ D) [PreservesLimitsOfSize.{v₃, u₃} F] : +lemma preservesSmallestLimits_of_preservesLimits (F : C ⥤ D) [PreservesLimitsOfSize.{v₃, u₃} F] : PreservesLimitsOfSize.{0, 0} F := - preservesLimitsOfSizeShrink F + preservesLimitsOfSize_shrink F + +@[deprecated "use preservesSmallestLimits_of_preservesLimits" (since := "2024-11-19")] +lemma preservesSmallestLimitsOfPreservesLimits (F : C ⥤ D) [PreservesLimitsOfSize.{v₃, u₃} F] : + PreservesLimitsOfSize.{0, 0} F := + preservesSmallestLimits_of_preservesLimits F /-- If F preserves one colimit cocone for the diagram K, then it preserves any colimit cocone for K. -/ -def preservesColimitOfPreservesColimitCocone {F : C ⥤ D} {t : Cocone K} (h : IsColimit t) +lemma preservesColimit_of_preserves_colimit_cocone {F : C ⥤ D} {t : Cocone K} (h : IsColimit t) (hF : IsColimit (F.mapCocone t)) : PreservesColimit K F := - ⟨fun h' => IsColimit.ofIsoColimit hF (Functor.mapIso _ (IsColimit.uniqueUpToIso h h'))⟩ + ⟨fun h' => ⟨IsColimit.ofIsoColimit hF (Functor.mapIso _ (IsColimit.uniqueUpToIso h h'))⟩⟩ + +@[deprecated "use preservesColimit_of_preserves_colimit_cocone" (since := "2024-11-19")] +lemma preservesColimitOfPreservesColimitCocone {F : C ⥤ D} {t : Cocone K} (h : IsColimit t) + (hF : IsColimit (F.mapCocone t)) : PreservesColimit K F := +preservesColimit_of_preserves_colimit_cocone h hF /-- Transfer preservation of colimits along a natural isomorphism in the shape. -/ -def preservesColimitOfIsoDiagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) [PreservesColimit K₁ F] : +lemma preservesColimit_of_iso_diagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) + [PreservesColimit K₁ F] : PreservesColimit K₂ F where - preserves {c} t := by + preserves {c} t := ⟨by apply IsColimit.precomposeHomEquiv (isoWhiskerRight h F : _) _ _ have := (IsColimit.precomposeHomEquiv h c).symm t apply IsColimit.ofIsoColimit (isColimitOfPreserves F this) - exact Cocones.ext (Iso.refl _) + exact Cocones.ext (Iso.refl _)⟩ + +@[deprecated "use preservesColimit_of_iso_diagram" (since := "2024-11-19")] +lemma preservesColimitOfIsoDiagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) + [PreservesColimit K₁ F] : + PreservesColimit K₂ F := + preservesColimit_of_iso_diagram F h /-- Transfer preservation of a colimit along a natural isomorphism in the functor. -/ -def preservesColimitOfNatIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [PreservesColimit K F] : +lemma preservesColimit_of_natIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [PreservesColimit K F] : PreservesColimit K G where - preserves t := IsColimit.mapCoconeEquiv h (PreservesColimit.preserves t) + preserves t := ⟨IsColimit.mapCoconeEquiv h (isColimitOfPreserves F t)⟩ + +@[deprecated preservesColimit_of_natIso (since := "2024-11-19")] +lemma preservesColimitOfNatIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [PreservesColimit K F] : + PreservesColimit K G := + preservesColimit_of_natIso K h /-- Transfer preservation of colimits of shape along a natural isomorphism in the functor. -/ -def preservesColimitsOfShapeOfNatIso {F G : C ⥤ D} (h : F ≅ G) [PreservesColimitsOfShape J F] : +lemma preservesColimitsOfShape_of_natIso {F G : C ⥤ D} (h : F ≅ G) [PreservesColimitsOfShape J F] : PreservesColimitsOfShape J G where - preservesColimit {K} := preservesColimitOfNatIso K h + preservesColimit {K} := preservesColimit_of_natIso K h + +@[deprecated "use preservesColimitsOfShape_of_natIso" (since := "2024-11-19")] +lemma preservesColimitsOfShapeOfNatIso {F G : C ⥤ D} (h : F ≅ G) [PreservesColimitsOfShape J F] : + PreservesColimitsOfShape J G := + preservesColimitsOfShape_of_natIso h /-- Transfer preservation of colimits along a natural isomorphism in the functor. -/ -def preservesColimitsOfNatIso {F G : C ⥤ D} (h : F ≅ G) [PreservesColimitsOfSize.{w, w'} F] : +lemma preservesColimits_of_natIso {F G : C ⥤ D} (h : F ≅ G) [PreservesColimitsOfSize.{w, w'} F] : PreservesColimitsOfSize.{w, w'} G where - preservesColimitsOfShape {_J} _𝒥₁ := preservesColimitsOfShapeOfNatIso h + preservesColimitsOfShape {_J} _𝒥₁ := preservesColimitsOfShape_of_natIso h + +@[deprecated "use preservesColimits_of_natIso" (since := "2024-11-19")] +lemma preservesColimitsOfNatIso {F G : C ⥤ D} (h : F ≅ G) [PreservesColimitsOfSize.{w, w'} F] : + PreservesColimitsOfSize.{w, w'} G := + preservesColimits_of_natIso h /-- Transfer preservation of colimits along an equivalence in the shape. -/ -def preservesColimitsOfShapeOfEquiv {J' : Type w₂} [Category.{w₂'} J'] (e : J ≌ J') (F : C ⥤ D) +lemma preservesColimitsOfShape_of_equiv {J' : Type w₂} [Category.{w₂'} J'] (e : J ≌ J') (F : C ⥤ D) [PreservesColimitsOfShape J F] : PreservesColimitsOfShape J' F where preservesColimit {K} := - { - preserves := fun {c} t => by + { preserves := fun {c} t => ⟨by let equ := e.invFunIdAssoc (K ⋙ F) have := (isColimitOfPreserves F (t.whiskerEquivalence e)).whiskerEquivalence e.symm apply ((IsColimit.precomposeInvEquiv equ _).symm this).ofIsoColimit refine Cocones.ext (Iso.refl _) fun j => ?_ dsimp - simp [equ, ← Functor.map_comp] } + simp [equ, ← Functor.map_comp]⟩ } + +@[deprecated "use preservesColimitsOfShape_of_equiv" (since := "2024-11-19")] +lemma preservesColimitsOfShapeOfEquiv {J' : Type w₂} [Category.{w₂'} J'] (e : J ≌ J') (F : C ⥤ D) + [PreservesColimitsOfShape J F] : PreservesColimitsOfShape J' F := + preservesColimitsOfShape_of_equiv e F /-- A functor preserving larger colimits also preserves smaller colimits. -/ -def preservesColimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] +lemma preservesColimitsOfSize_of_univLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] [PreservesColimitsOfSize.{w', w₂'} F] : PreservesColimitsOfSize.{w, w₂} F where - preservesColimitsOfShape {J} := preservesColimitsOfShapeOfEquiv + preservesColimitsOfShape {J} := preservesColimitsOfShape_of_equiv ((ShrinkHoms.equivalence J).trans <| Shrink.equivalence _).symm F +@[deprecated "use preservesColimitsOfSize_of_univLE" (since := "2024-11-19")] +lemma preservesColimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] + [PreservesColimitsOfSize.{w', w₂'} F] : PreservesColimitsOfSize.{w, w₂} F := + preservesColimitsOfSize_of_univLE.{w', w₂'} F + -- See library note [dsimp, simp]. /-- -`PreservesColimitsOfSizeShrink.{w w'} F` tries to obtain `PreservesColimitsOfSize.{w w'} F` +`PreservesColimitsOfSize_shrink.{w w'} F` tries to obtain `PreservesColimitsOfSize.{w w'} F` from some other `PreservesColimitsOfSize F`. -/ -def preservesColimitsOfSizeShrink (F : C ⥤ D) [PreservesColimitsOfSize.{max w w₂, max w' w₂'} F] : - PreservesColimitsOfSize.{w, w'} F := preservesColimitsOfSizeOfUnivLE.{max w w₂, max w' w₂'} F +lemma preservesColimitsOfSize_shrink (F : C ⥤ D) + [PreservesColimitsOfSize.{max w w₂, max w' w₂'} F] : + PreservesColimitsOfSize.{w, w'} F := preservesColimitsOfSize_of_univLE.{max w w₂, max w' w₂'} F + +@[deprecated "use preservesColimitsOfSize_shrink" (since := "2024-11-19")] +lemma PreservesColimitsOfSizeShrink (F : C ⥤ D) + [PreservesColimitsOfSize.{max w w₂, max w' w₂'} F] : + PreservesColimitsOfSize.{w, w'} F := + preservesColimitsOfSize_shrink F /-- Preserving colimits at any universe implies preserving colimits at universe `0`. -/ -def preservesSmallestColimitsOfPreservesColimits (F : C ⥤ D) [PreservesColimitsOfSize.{v₃, u₃} F] : +lemma preservesSmallestColimits_of_preservesColimits (F : C ⥤ D) + [PreservesColimitsOfSize.{v₃, u₃} F] : PreservesColimitsOfSize.{0, 0} F := - preservesColimitsOfSizeShrink F + preservesColimitsOfSize_shrink F + +@[deprecated "use preservesSmallestColimits_of_preservesColimits" (since := "2024-11-19")] +lemma preservesSmallestColimitsOfPreservesColimits (F : C ⥤ D) + [PreservesColimitsOfSize.{v₃, u₃} F] : + PreservesColimitsOfSize.{0, 0} F := + preservesSmallestColimits_of_preservesColimits F /-- A functor `F : C ⥤ D` reflects limits for `K : J ⥤ C` if whenever the image of a cone over `K` under `F` is a limit cone in `D`, the cone was already a limit cone in `C`. Note that we do not assume a priori that `D` actually has any limits. -/ -class ReflectsLimit (K : J ⥤ C) (F : C ⥤ D) where - reflects : ∀ {c : Cone K}, IsLimit (F.mapCone c) → IsLimit c +class ReflectsLimit (K : J ⥤ C) (F : C ⥤ D) : Prop where + reflects {c : Cone K} (hc : IsLimit (F.mapCone c)) : Nonempty (IsLimit c) /-- A functor `F : C ⥤ D` reflects colimits for `K : J ⥤ C` if whenever the image of a cocone over `K` under `F` is a colimit cocone in `D`, the cocone was already a colimit cocone in `C`. Note that we do not assume a priori that `D` actually has any colimits. -/ -class ReflectsColimit (K : J ⥤ C) (F : C ⥤ D) where - reflects : ∀ {c : Cocone K}, IsColimit (F.mapCocone c) → IsColimit c +class ReflectsColimit (K : J ⥤ C) (F : C ⥤ D) : Prop where + reflects {c : Cocone K} (hc : IsColimit (F.mapCocone c)) : Nonempty (IsColimit c) /-- A functor `F : C ⥤ D` reflects limits of shape `J` if whenever the image of a cone over some `K : J ⥤ C` under `F` is a limit cone in `D`, the cone was already a limit cone in `C`. Note that we do not assume a priori that `D` actually has any limits. -/ -class ReflectsLimitsOfShape (J : Type w) [Category.{w'} J] (F : C ⥤ D) where +class ReflectsLimitsOfShape (J : Type w) [Category.{w'} J] (F : C ⥤ D) : Prop where reflectsLimit : ∀ {K : J ⥤ C}, ReflectsLimit K F := by infer_instance /-- A functor `F : C ⥤ D` reflects colimits of shape `J` if @@ -348,7 +479,7 @@ whenever the image of a cocone over some `K : J ⥤ C` under `F` is a colimit co the cocone was already a colimit cocone in `C`. Note that we do not assume a priori that `D` actually has any colimits. -/ -class ReflectsColimitsOfShape (J : Type w) [Category.{w'} J] (F : C ⥤ D) where +class ReflectsColimitsOfShape (J : Type w) [Category.{w'} J] (F : C ⥤ D) : Prop where reflectsColimit : ∀ {K : J ⥤ C}, ReflectsColimit K F := by infer_instance -- This should be used with explicit universe variables. @@ -358,7 +489,7 @@ the cone was already a limit cone in `C`. Note that we do not assume a priori that `D` actually has any limits. -/ @[nolint checkUnivs, pp_with_univ] -class ReflectsLimitsOfSize (F : C ⥤ D) where +class ReflectsLimitsOfSize (F : C ⥤ D) : Prop where reflectsLimitsOfShape : ∀ {J : Type w} [Category.{w'} J], ReflectsLimitsOfShape J F := by infer_instance @@ -377,7 +508,7 @@ the cocone was already a colimit cocone in `C`. Note that we do not assume a priori that `D` actually has any colimits. -/ @[nolint checkUnivs, pp_with_univ] -class ReflectsColimitsOfSize (F : C ⥤ D) where +class ReflectsColimitsOfSize (F : C ⥤ D) : Prop where reflectsColimitsOfShape : ∀ {J : Type w} [Category.{w'} J], ReflectsColimitsOfShape J F := by infer_instance @@ -393,7 +524,8 @@ abbrev ReflectsColimits (F : C ⥤ D) := guide typeclass resolution. -/ def isLimitOfReflects (F : C ⥤ D) {c : Cone K} (t : IsLimit (F.mapCone c)) - [ReflectsLimit K F] : IsLimit c := ReflectsLimit.reflects t + [ReflectsLimit K F] : IsLimit c := + (ReflectsLimit.reflects t).some /-- A convenience function for `ReflectsColimit`, which takes the functor as an explicit argument to @@ -401,176 +533,259 @@ guide typeclass resolution. -/ def isColimitOfReflects (F : C ⥤ D) {c : Cocone K} (t : IsColimit (F.mapCocone c)) [ReflectsColimit K F] : IsColimit c := - ReflectsColimit.reflects t + (ReflectsColimit.reflects t).some instance reflectsLimit_subsingleton (K : J ⥤ C) (F : C ⥤ D) : Subsingleton (ReflectsLimit K F) := by - constructor; rintro ⟨a⟩ ⟨b⟩; congr!; subsingleton + constructor; rintro ⟨a⟩ ⟨b⟩; congr! instance reflectsColimit_subsingleton (K : J ⥤ C) (F : C ⥤ D) : Subsingleton (ReflectsColimit K F) := by - constructor; rintro ⟨a⟩ ⟨b⟩; congr!; subsingleton + constructor; rintro ⟨a⟩ ⟨b⟩; congr! instance reflectsLimitsOfShape_subsingleton (J : Type w) [Category.{w'} J] (F : C ⥤ D) : Subsingleton (ReflectsLimitsOfShape J F) := by - constructor; rintro ⟨a⟩ ⟨b⟩; congr!; subsingleton + constructor; rintro ⟨a⟩ ⟨b⟩; congr! instance reflectsColimitsOfShape_subsingleton (J : Type w) [Category.{w'} J] (F : C ⥤ D) : Subsingleton (ReflectsColimitsOfShape J F) := by - constructor; rintro ⟨a⟩ ⟨b⟩; congr!; subsingleton + constructor; rintro ⟨a⟩ ⟨b⟩; congr! instance reflects_limits_subsingleton (F : C ⥤ D) : Subsingleton (ReflectsLimitsOfSize.{w', w} F) := by - constructor; rintro ⟨a⟩ ⟨b⟩; congr!; subsingleton + constructor; rintro ⟨a⟩ ⟨b⟩; congr! instance reflects_colimits_subsingleton (F : C ⥤ D) : Subsingleton (ReflectsColimitsOfSize.{w', w} F) := by - constructor; rintro ⟨a⟩ ⟨b⟩; congr!; subsingleton + constructor; rintro ⟨a⟩ ⟨b⟩; congr! -- see Note [lower instance priority] -instance (priority := 100) reflectsLimitOfReflectsLimitsOfShape (K : J ⥤ C) (F : C ⥤ D) +instance (priority := 100) reflectsLimit_of_reflectsLimitsOfShape (K : J ⥤ C) (F : C ⥤ D) [ReflectsLimitsOfShape J F] : ReflectsLimit K F := ReflectsLimitsOfShape.reflectsLimit -- see Note [lower instance priority] -instance (priority := 100) reflectsColimitOfReflectsColimitsOfShape (K : J ⥤ C) (F : C ⥤ D) +instance (priority := 100) reflectsColimit_of_reflectsColimitsOfShape (K : J ⥤ C) (F : C ⥤ D) [ReflectsColimitsOfShape J F] : ReflectsColimit K F := ReflectsColimitsOfShape.reflectsColimit -- see Note [lower instance priority] -instance (priority := 100) reflectsLimitsOfShapeOfReflectsLimits (J : Type w) [Category.{w'} J] +instance (priority := 100) reflectsLimitsOfShape_of_reflectsLimits (J : Type w) [Category.{w'} J] (F : C ⥤ D) [ReflectsLimitsOfSize.{w', w} F] : ReflectsLimitsOfShape J F := ReflectsLimitsOfSize.reflectsLimitsOfShape -- see Note [lower instance priority] -instance (priority := 100) reflectsColimitsOfShapeOfReflectsColimits (J : Type w) [Category.{w'} J] +instance (priority := 100) reflectsColimitsOfShape_of_reflectsColimits + (J : Type w) [Category.{w'} J] (F : C ⥤ D) [ReflectsColimitsOfSize.{w', w} F] : ReflectsColimitsOfShape J F := ReflectsColimitsOfSize.reflectsColimitsOfShape -instance idReflectsLimits : ReflectsLimitsOfSize.{w, w'} (𝟭 C) where +instance id_reflectsLimits : ReflectsLimitsOfSize.{w, w'} (𝟭 C) where reflectsLimitsOfShape {J} 𝒥 := - { - reflectsLimit := fun {K} => + { reflectsLimit := fun {K} => ⟨fun {c} h => ⟨fun s => h.lift ⟨s.pt, fun j => s.π.app j, fun _ _ f => s.π.naturality f⟩, by cases K; rcases c with ⟨_, _, _⟩; intro s j; cases s; exact h.fac _ j, by cases K; rcases c with ⟨_, _, _⟩; intro s m w; rcases s with ⟨_, _, _⟩; exact h.uniq _ m w⟩⟩ } -instance idReflectsColimits : ReflectsColimitsOfSize.{w, w'} (𝟭 C) where +@[deprecated "use id_reflectsLimits" (since := "2024-11-19")] +lemma idReflectsLimits : ReflectsLimitsOfSize.{w, w'} (𝟭 C) := id_reflectsLimits + +instance id_reflectsColimits : ReflectsColimitsOfSize.{w, w'} (𝟭 C) where reflectsColimitsOfShape {J} 𝒥 := - { - reflectsColimit := fun {K} => + { reflectsColimit := fun {K} => ⟨fun {c} h => ⟨fun s => h.desc ⟨s.pt, fun j => s.ι.app j, fun _ _ f => s.ι.naturality f⟩, by cases K; rcases c with ⟨_, _, _⟩; intro s j; cases s; exact h.fac _ j, by cases K; rcases c with ⟨_, _, _⟩; intro s m w; rcases s with ⟨_, _, _⟩; exact h.uniq _ m w⟩⟩ } +@[deprecated "use id_reflectsColimits" (since := "2024-11-19")] +lemma idReflectsColimits : ReflectsColimitsOfSize.{w, w'} (𝟭 C) := id_reflectsColimits + section variable {E : Type u₃} [ℰ : Category.{v₃} E] variable (F : C ⥤ D) (G : D ⥤ E) -instance compReflectsLimit [ReflectsLimit K F] [ReflectsLimit (K ⋙ F) G] : +instance comp_reflectsLimit [ReflectsLimit K F] [ReflectsLimit (K ⋙ F) G] : ReflectsLimit K (F ⋙ G) := - ⟨fun h => ReflectsLimit.reflects (ReflectsLimit.reflects h)⟩ + ⟨fun h => ReflectsLimit.reflects (isLimitOfReflects G h)⟩ -instance compReflectsLimitsOfShape [ReflectsLimitsOfShape J F] [ReflectsLimitsOfShape J G] : +instance comp_reflectsLimitsOfShape [ReflectsLimitsOfShape J F] [ReflectsLimitsOfShape J G] : ReflectsLimitsOfShape J (F ⋙ G) where -instance compReflectsLimits [ReflectsLimitsOfSize.{w', w} F] [ReflectsLimitsOfSize.{w', w} G] : +instance comp_reflectsLimits [ReflectsLimitsOfSize.{w', w} F] [ReflectsLimitsOfSize.{w', w} G] : ReflectsLimitsOfSize.{w', w} (F ⋙ G) where -instance compReflectsColimit [ReflectsColimit K F] [ReflectsColimit (K ⋙ F) G] : +instance comp_reflectsColimit [ReflectsColimit K F] [ReflectsColimit (K ⋙ F) G] : ReflectsColimit K (F ⋙ G) := - ⟨fun h => ReflectsColimit.reflects (ReflectsColimit.reflects h)⟩ + ⟨fun h => ReflectsColimit.reflects (isColimitOfReflects G h)⟩ -instance compReflectsColimitsOfShape [ReflectsColimitsOfShape J F] [ReflectsColimitsOfShape J G] : +instance comp_reflectsColimitsOfShape [ReflectsColimitsOfShape J F] [ReflectsColimitsOfShape J G] : ReflectsColimitsOfShape J (F ⋙ G) where -instance compReflectsColimits [ReflectsColimitsOfSize.{w', w} F] +instance comp_reflectsColimits [ReflectsColimitsOfSize.{w', w} F] [ReflectsColimitsOfSize.{w', w} G] : ReflectsColimitsOfSize.{w', w} (F ⋙ G) where +@[deprecated "use comp_reflectsLimit" (since := "2024-11-19")] +lemma compReflectsLimit [ReflectsLimit K F] [ReflectsLimit (K ⋙ F) G] : + ReflectsLimit K (F ⋙ G) := inferInstance + +@[deprecated "use comp_reflectsLimitsOfShape " (since := "2024-11-19")] +lemma compReflectsLimitsOfShape [ReflectsLimitsOfShape J F] [ReflectsLimitsOfShape J G] : + ReflectsLimitsOfShape J (F ⋙ G) := inferInstance + +@[deprecated "use comp_reflectsLimits" (since := "2024-11-19")] +lemma compReflectsLimits [ReflectsLimitsOfSize.{w', w} F] [ReflectsLimitsOfSize.{w', w} G] : + ReflectsLimitsOfSize.{w', w} (F ⋙ G) := inferInstance + +@[deprecated "use comp_reflectsColimit" (since := "2024-11-19")] +lemma compReflectsColimit [ReflectsColimit K F] [ReflectsColimit (K ⋙ F) G] : + ReflectsColimit K (F ⋙ G) := inferInstance + +@[deprecated "use comp_reflectsColimitsOfShape " (since := "2024-11-19")] +lemma compReflectsColimitsOfShape [ReflectsColimitsOfShape J F] [ReflectsColimitsOfShape J G] : + ReflectsColimitsOfShape J (F ⋙ G) := inferInstance + +@[deprecated "use comp_reflectsColimits" (since := "2024-11-19")] +lemma compReflectsColimits [ReflectsColimitsOfSize.{w', w} F] [ReflectsColimitsOfSize.{w', w} G] : + ReflectsColimitsOfSize.{w', w} (F ⋙ G) := inferInstance + /-- If `F ⋙ G` preserves limits for `K`, and `G` reflects limits for `K ⋙ F`, then `F` preserves limits for `K`. -/ -def preservesLimitOfReflectsOfPreserves [PreservesLimit K (F ⋙ G)] [ReflectsLimit (K ⋙ F) G] : +lemma preservesLimit_of_reflects_of_preserves [PreservesLimit K (F ⋙ G)] [ReflectsLimit (K ⋙ F) G] : PreservesLimit K F := - ⟨fun h => by + ⟨fun h => ⟨by apply isLimitOfReflects G - apply isLimitOfPreserves (F ⋙ G) h⟩ + apply isLimitOfPreserves (F ⋙ G) h⟩⟩ + +@[deprecated "use preservesLimit_of_reflects_of_preserves" (since := "2024-11-19")] +lemma preservesLimitOfReflectsOfPreserves [PreservesLimit K (F ⋙ G)] [ReflectsLimit (K ⋙ F) G] : + PreservesLimit K F := + preservesLimit_of_reflects_of_preserves F G /-- If `F ⋙ G` preserves limits of shape `J` and `G` reflects limits of shape `J`, then `F` preserves limits of shape `J`. -/ -def preservesLimitsOfShapeOfReflectsOfPreserves [PreservesLimitsOfShape J (F ⋙ G)] +lemma preservesLimitsOfShape_of_reflects_of_preserves [PreservesLimitsOfShape J (F ⋙ G)] [ReflectsLimitsOfShape J G] : PreservesLimitsOfShape J F where - preservesLimit := preservesLimitOfReflectsOfPreserves F G + preservesLimit := preservesLimit_of_reflects_of_preserves F G + +@[deprecated "use preservesLimitsOfShape_of_reflects_of_preserves" (since := "2024-11-19")] +lemma preservesLimitsOfShapeOfReflectsOfPreserves [PreservesLimitsOfShape J (F ⋙ G)] + [ReflectsLimitsOfShape J G] : PreservesLimitsOfShape J F := + preservesLimitsOfShape_of_reflects_of_preserves F G /-- If `F ⋙ G` preserves limits and `G` reflects limits, then `F` preserves limits. -/ -def preservesLimitsOfReflectsOfPreserves [PreservesLimitsOfSize.{w', w} (F ⋙ G)] +lemma preservesLimits_of_reflects_of_preserves [PreservesLimitsOfSize.{w', w} (F ⋙ G)] [ReflectsLimitsOfSize.{w', w} G] : PreservesLimitsOfSize.{w', w} F where - preservesLimitsOfShape := preservesLimitsOfShapeOfReflectsOfPreserves F G + preservesLimitsOfShape := preservesLimitsOfShape_of_reflects_of_preserves F G + +@[deprecated "use preservesLimits_of_reflects_of_preserves" (since := "2024-11-19")] +lemma preservesLimitsOfReflectsOfPreserves [PreservesLimitsOfSize.{w', w} (F ⋙ G)] + [ReflectsLimitsOfSize.{w', w} G] : PreservesLimitsOfSize.{w', w} F := + preservesLimits_of_reflects_of_preserves.{w', w} F G /-- Transfer reflection of limits along a natural isomorphism in the diagram. -/ -def reflectsLimitOfIsoDiagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) [ReflectsLimit K₁ F] : +lemma reflectsLimit_of_iso_diagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) [ReflectsLimit K₁ F] : ReflectsLimit K₂ F where - reflects {c} t := by + reflects {c} t := ⟨by apply IsLimit.postcomposeInvEquiv h c (isLimitOfReflects F _) apply ((IsLimit.postcomposeInvEquiv (isoWhiskerRight h F : _) _).symm t).ofIsoLimit _ - exact Cones.ext (Iso.refl _) + exact Cones.ext (Iso.refl _)⟩ + +@[deprecated "use reflectsLimit_of_iso_diagram" (since := "2024-11-19")] +lemma reflectsLimitOfIsoDiagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) [ReflectsLimit K₁ F] : + ReflectsLimit K₂ F := + reflectsLimit_of_iso_diagram F h /-- Transfer reflection of a limit along a natural isomorphism in the functor. -/ -def reflectsLimitOfNatIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [ReflectsLimit K F] : +lemma reflectsLimit_of_natIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [ReflectsLimit K F] : ReflectsLimit K G where reflects t := ReflectsLimit.reflects (IsLimit.mapConeEquiv h.symm t) +@[deprecated "use reflectsLimit_of_natIso" (since := "2024-11-19")] +lemma reflectsLimitOfNatIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [ReflectsLimit K F] : + ReflectsLimit K G := + reflectsLimit_of_natIso K h + /-- Transfer reflection of limits of shape along a natural isomorphism in the functor. -/ -def reflectsLimitsOfShapeOfNatIso {F G : C ⥤ D} (h : F ≅ G) [ReflectsLimitsOfShape J F] : +lemma reflectsLimitsOfShape_of_natIso {F G : C ⥤ D} (h : F ≅ G) [ReflectsLimitsOfShape J F] : ReflectsLimitsOfShape J G where - reflectsLimit {K} := reflectsLimitOfNatIso K h + reflectsLimit {K} := reflectsLimit_of_natIso K h + +@[deprecated "use reflectsLimitsOfShape_of_natIso" (since := "2024-11-19")] +lemma reflectsLimitsOfShapeOfNatIso {F G : C ⥤ D} (h : F ≅ G) [ReflectsLimitsOfShape J F] : + ReflectsLimitsOfShape J G := + reflectsLimitsOfShape_of_natIso h /-- Transfer reflection of limits along a natural isomorphism in the functor. -/ -def reflectsLimitsOfNatIso {F G : C ⥤ D} (h : F ≅ G) [ReflectsLimitsOfSize.{w', w} F] : +lemma reflectsLimits_of_natIso {F G : C ⥤ D} (h : F ≅ G) [ReflectsLimitsOfSize.{w', w} F] : ReflectsLimitsOfSize.{w', w} G where - reflectsLimitsOfShape := reflectsLimitsOfShapeOfNatIso h + reflectsLimitsOfShape := reflectsLimitsOfShape_of_natIso h + +@[deprecated "use reflectsLimits_of_natIso" (since := "2024-11-19")] +lemma reflectsLimitsOfNatIso {F G : C ⥤ D} (h : F ≅ G) [ReflectsLimitsOfSize.{w', w} F] : + ReflectsLimitsOfSize.{w', w} G := + reflectsLimits_of_natIso h /-- Transfer reflection of limits along an equivalence in the shape. -/ -def reflectsLimitsOfShapeOfEquiv {J' : Type w₂} [Category.{w₂'} J'] (e : J ≌ J') (F : C ⥤ D) +lemma reflectsLimitsOfShape_of_equiv {J' : Type w₂} [Category.{w₂'} J'] (e : J ≌ J') (F : C ⥤ D) [ReflectsLimitsOfShape J F] : ReflectsLimitsOfShape J' F where reflectsLimit {K} := - { - reflects := fun {c} t => by + { reflects := fun {c} t => ⟨by apply IsLimit.ofWhiskerEquivalence e apply isLimitOfReflects F apply IsLimit.ofIsoLimit _ (Functor.mapConeWhisker _).symm - exact IsLimit.whiskerEquivalence t _ } + exact IsLimit.whiskerEquivalence t _⟩ } + +@[deprecated "use reflectsLimitsOfShape_of_equiv" (since := "2024-11-19")] +lemma reflectsLimitsOfShapeOfEquiv {J' : Type w₂} [Category.{w₂'} J'] (e : J ≌ J') (F : C ⥤ D) + [ReflectsLimitsOfShape J F] : ReflectsLimitsOfShape J' F := + reflectsLimitsOfShape_of_equiv e F /-- A functor reflecting larger limits also reflects smaller limits. -/ -def reflectsLimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] +lemma reflectsLimitsOfSize_of_univLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] [ReflectsLimitsOfSize.{w', w₂'} F] : ReflectsLimitsOfSize.{w, w₂} F where - reflectsLimitsOfShape {J} := reflectsLimitsOfShapeOfEquiv + reflectsLimitsOfShape {J} := reflectsLimitsOfShape_of_equiv ((ShrinkHoms.equivalence J).trans <| Shrink.equivalence _).symm F -/-- `reflectsLimitsOfSizeShrink.{w w'} F` tries to obtain `reflectsLimitsOfSize.{w w'} F` +@[deprecated "use reflectsLimitsOfSize_of_univLE" (since := "2024-11-19")] +lemma reflectsLimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] + [ReflectsLimitsOfSize.{w', w₂'} F] : ReflectsLimitsOfSize.{w, w₂} F := + reflectsLimitsOfSize_of_univLE.{w'} F + +/-- `reflectsLimitsOfSize_shrink.{w w'} F` tries to obtain `reflectsLimitsOfSize.{w w'} F` from some other `reflectsLimitsOfSize F`. -/ -def reflectsLimitsOfSizeShrink (F : C ⥤ D) [ReflectsLimitsOfSize.{max w w₂, max w' w₂'} F] : - ReflectsLimitsOfSize.{w, w'} F := reflectsLimitsOfSizeOfUnivLE.{max w w₂, max w' w₂'} F +lemma reflectsLimitsOfSize_shrink (F : C ⥤ D) [ReflectsLimitsOfSize.{max w w₂, max w' w₂'} F] : + ReflectsLimitsOfSize.{w, w'} F := reflectsLimitsOfSize_of_univLE.{max w w₂, max w' w₂'} F + +@[deprecated "use reflectsLimitsOfSize_shrink" (since := "2024-11-19")] +lemma reflectsLimitsOfSizeShrink (F : C ⥤ D) [ReflectsLimitsOfSize.{max w w₂, max w' w₂'} F] : + ReflectsLimitsOfSize.{w, w'} F := + reflectsLimitsOfSize_shrink F /-- Reflecting limits at any universe implies reflecting limits at universe `0`. -/ -def reflectsSmallestLimitsOfReflectsLimits (F : C ⥤ D) [ReflectsLimitsOfSize.{v₃, u₃} F] : +lemma reflectsSmallestLimits_of_reflectsLimits (F : C ⥤ D) [ReflectsLimitsOfSize.{v₃, u₃} F] : + ReflectsLimitsOfSize.{0, 0} F := + reflectsLimitsOfSize_shrink F + +@[deprecated "use reflectsSmallestLimits_of_reflectsLimits" (since := "2024-11-19")] +lemma reflectsSmallestLimitsOfReflectsLimits (F : C ⥤ D) [ReflectsLimitsOfSize.{v₃, u₃} F] : ReflectsLimitsOfSize.{0, 0} F := - reflectsLimitsOfSizeShrink F + reflectsSmallestLimits_of_reflectsLimits F /-- If the limit of `F` exists and `G` preserves it, then if `G` reflects isomorphisms then it reflects the limit of `F`. -/ -- Porting note: previous behavior of apply pushed instance holes into hypotheses, this errors -def reflectsLimitOfReflectsIsomorphisms (F : J ⥤ C) (G : C ⥤ D) [G.ReflectsIsomorphisms] +lemma reflectsLimit_of_reflectsIsomorphisms (F : J ⥤ C) (G : C ⥤ D) [G.ReflectsIsomorphisms] [HasLimit F] [PreservesLimit F G] : ReflectsLimit F G where reflects {c} t := by - suffices IsIso (IsLimit.lift (limit.isLimit F) c) from by - apply IsLimit.ofPointIso (limit.isLimit F) + suffices IsIso (IsLimit.lift (limit.isLimit F) c) from ⟨by + apply IsLimit.ofPointIso (limit.isLimit F)⟩ change IsIso ((Cones.forget _).map ((limit.isLimit F).liftConeMorphism c)) suffices IsIso (IsLimit.liftConeMorphism (limit.isLimit F) c) from by apply (Cones.forget F).map_isIso _ @@ -579,100 +794,174 @@ def reflectsLimitOfReflectsIsomorphisms (F : J ⥤ C) (G : C ⥤ D) [G.ReflectsI apply isIso_of_reflects_iso _ (Cones.functoriality F G) exact t.hom_isIso (isLimitOfPreserves G (limit.isLimit F)) _ +@[deprecated "use reflectsLimit_of_reflectsIsomorphisms" (since := "2024-11-19")] +lemma reflectsLimitOfReflectsIsomorphisms (F : J ⥤ C) (G : C ⥤ D) [G.ReflectsIsomorphisms] + [HasLimit F] [PreservesLimit F G] : ReflectsLimit F G := + reflectsLimit_of_reflectsIsomorphisms F G + /-- If `C` has limits of shape `J` and `G` preserves them, then if `G` reflects isomorphisms then it reflects limits of shape `J`. -/ -def reflectsLimitsOfShapeOfReflectsIsomorphisms {G : C ⥤ D} [G.ReflectsIsomorphisms] +lemma reflectsLimitsOfShape_of_reflectsIsomorphisms {G : C ⥤ D} [G.ReflectsIsomorphisms] [HasLimitsOfShape J C] [PreservesLimitsOfShape J G] : ReflectsLimitsOfShape J G where - reflectsLimit {F} := reflectsLimitOfReflectsIsomorphisms F G + reflectsLimit {F} := reflectsLimit_of_reflectsIsomorphisms F G + +@[deprecated "use reflectsLimitsOfShape_of_reflectsIsomorphisms" (since := "2024-11-19")] +lemma reflectsLimitsOfShapeOfReflectsIsomorphisms {G : C ⥤ D} [G.ReflectsIsomorphisms] + [HasLimitsOfShape J C] [PreservesLimitsOfShape J G] : ReflectsLimitsOfShape J G := + reflectsLimitsOfShape_of_reflectsIsomorphisms /-- If `C` has limits and `G` preserves limits, then if `G` reflects isomorphisms then it reflects limits. -/ -def reflectsLimitsOfReflectsIsomorphisms {G : C ⥤ D} [G.ReflectsIsomorphisms] +lemma reflectsLimits_of_reflectsIsomorphisms {G : C ⥤ D} [G.ReflectsIsomorphisms] [HasLimitsOfSize.{w', w} C] [PreservesLimitsOfSize.{w', w} G] : ReflectsLimitsOfSize.{w', w} G where - reflectsLimitsOfShape := reflectsLimitsOfShapeOfReflectsIsomorphisms + reflectsLimitsOfShape := reflectsLimitsOfShape_of_reflectsIsomorphisms + +@[deprecated "use reflectsLimits_of_reflectsIsomorphisms" (since := "2024-11-19")] +lemma reflectsLimitsOfReflectsIsomorphisms {G : C ⥤ D} [G.ReflectsIsomorphisms] + [HasLimitsOfSize.{w', w} C] [PreservesLimitsOfSize.{w', w} G] : + ReflectsLimitsOfSize.{w', w} G := + reflectsLimits_of_reflectsIsomorphisms /-- If `F ⋙ G` preserves colimits for `K`, and `G` reflects colimits for `K ⋙ F`, then `F` preserves colimits for `K`. -/ -def preservesColimitOfReflectsOfPreserves [PreservesColimit K (F ⋙ G)] [ReflectsColimit (K ⋙ F) G] : +lemma preservesColimit_of_reflects_of_preserves + [PreservesColimit K (F ⋙ G)] [ReflectsColimit (K ⋙ F) G] : PreservesColimit K F := - ⟨fun {c} h => by + ⟨fun {c} h => ⟨by apply isColimitOfReflects G - apply isColimitOfPreserves (F ⋙ G) h⟩ + apply isColimitOfPreserves (F ⋙ G) h⟩⟩ + +@[deprecated "use preservesColimit_of_reflects_of_preserves" (since := "2024-11-19")] +lemma preservesColimitOfReflectsOfPreserves + [PreservesColimit K (F ⋙ G)] [ReflectsColimit (K ⋙ F) G] : + PreservesColimit K F := + preservesColimit_of_reflects_of_preserves F G /-- If `F ⋙ G` preserves colimits of shape `J` and `G` reflects colimits of shape `J`, then `F` preserves colimits of shape `J`. -/ -def preservesColimitsOfShapeOfReflectsOfPreserves [PreservesColimitsOfShape J (F ⋙ G)] +lemma preservesColimitsOfShape_of_reflects_of_preserves [PreservesColimitsOfShape J (F ⋙ G)] [ReflectsColimitsOfShape J G] : PreservesColimitsOfShape J F where - preservesColimit := preservesColimitOfReflectsOfPreserves F G + preservesColimit := preservesColimit_of_reflects_of_preserves F G + +@[deprecated "use preservesColimitsOfShape_of_reflects_of_preserves" (since := "2024-11-19")] +lemma preservesColimitsOfShapeOfReflectsOfPreserves [PreservesColimitsOfShape J (F ⋙ G)] + [ReflectsColimitsOfShape J G] : PreservesColimitsOfShape J F := + preservesColimitsOfShape_of_reflects_of_preserves F G /-- If `F ⋙ G` preserves colimits and `G` reflects colimits, then `F` preserves colimits. -/ -def preservesColimitsOfReflectsOfPreserves [PreservesColimitsOfSize.{w', w} (F ⋙ G)] +lemma preservesColimits_of_reflects_of_preserves [PreservesColimitsOfSize.{w', w} (F ⋙ G)] [ReflectsColimitsOfSize.{w', w} G] : PreservesColimitsOfSize.{w', w} F where - preservesColimitsOfShape := preservesColimitsOfShapeOfReflectsOfPreserves F G + preservesColimitsOfShape := preservesColimitsOfShape_of_reflects_of_preserves F G + +@[deprecated "use preservesColimits_of_reflects_of_preserves" (since := "2024-11-19")] +lemma preservesColimitsOfReflectsOfPreserves [PreservesColimitsOfSize.{w', w} (F ⋙ G)] + [ReflectsColimitsOfSize.{w', w} G] : PreservesColimitsOfSize.{w', w} F := + preservesColimits_of_reflects_of_preserves F G /-- Transfer reflection of colimits along a natural isomorphism in the diagram. -/ -def reflectsColimitOfIsoDiagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) [ReflectsColimit K₁ F] : +lemma reflectsColimit_of_iso_diagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) + [ReflectsColimit K₁ F] : ReflectsColimit K₂ F where - reflects {c} t := by + reflects {c} t := ⟨by apply IsColimit.precomposeHomEquiv h c (isColimitOfReflects F _) apply ((IsColimit.precomposeHomEquiv (isoWhiskerRight h F : _) _).symm t).ofIsoColimit _ - exact Cocones.ext (Iso.refl _) + exact Cocones.ext (Iso.refl _)⟩ + +@[deprecated "use reflectsColimit_of_iso_diagram" (since := "2024-11-19")] +lemma reflectsColimitOfIsoDiagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) + [ReflectsColimit K₁ F] : + ReflectsColimit K₂ F := + reflectsColimit_of_iso_diagram F h /-- Transfer reflection of a colimit along a natural isomorphism in the functor. -/ -def reflectsColimitOfNatIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [ReflectsColimit K F] : +lemma reflectsColimit_of_natIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [ReflectsColimit K F] : ReflectsColimit K G where reflects t := ReflectsColimit.reflects (IsColimit.mapCoconeEquiv h.symm t) +@[deprecated "use reflectsColimit_of_natIso" (since := "2024-11-19")] +lemma reflectsColimitOfNatIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [ReflectsColimit K F] : + ReflectsColimit K G := + reflectsColimit_of_natIso K h + /-- Transfer reflection of colimits of shape along a natural isomorphism in the functor. -/ -def reflectsColimitsOfShapeOfNatIso {F G : C ⥤ D} (h : F ≅ G) [ReflectsColimitsOfShape J F] : +lemma reflectsColimitsOfShape_of_natIso {F G : C ⥤ D} (h : F ≅ G) [ReflectsColimitsOfShape J F] : ReflectsColimitsOfShape J G where - reflectsColimit {K} := reflectsColimitOfNatIso K h + reflectsColimit {K} := reflectsColimit_of_natIso K h + +@[deprecated "use reflectsColimitsOfShape_of_natIso" (since := "2024-11-19")] +lemma reflectsColimitsOfShapeOfNatIso {F G : C ⥤ D} (h : F ≅ G) [ReflectsColimitsOfShape J F] : + ReflectsColimitsOfShape J G := + reflectsColimitsOfShape_of_natIso h /-- Transfer reflection of colimits along a natural isomorphism in the functor. -/ -def reflectsColimitsOfNatIso {F G : C ⥤ D} (h : F ≅ G) [ReflectsColimitsOfSize.{w, w'} F] : +lemma reflectsColimits_of_natIso {F G : C ⥤ D} (h : F ≅ G) [ReflectsColimitsOfSize.{w, w'} F] : ReflectsColimitsOfSize.{w, w'} G where - reflectsColimitsOfShape := reflectsColimitsOfShapeOfNatIso h + reflectsColimitsOfShape := reflectsColimitsOfShape_of_natIso h + +@[deprecated "use reflectsColimits_of_natIso" (since := "2024-11-19")] +lemma reflectsColimitsOfNatIso {F G : C ⥤ D} (h : F ≅ G) [ReflectsColimitsOfSize.{w, w'} F] : + ReflectsColimitsOfSize.{w, w'} G := + reflectsColimits_of_natIso h /-- Transfer reflection of colimits along an equivalence in the shape. -/ -def reflectsColimitsOfShapeOfEquiv {J' : Type w₂} [Category.{w₂'} J'] (e : J ≌ J') (F : C ⥤ D) +lemma reflectsColimitsOfShape_of_equiv {J' : Type w₂} [Category.{w₂'} J'] (e : J ≌ J') (F : C ⥤ D) [ReflectsColimitsOfShape J F] : ReflectsColimitsOfShape J' F where reflectsColimit := - { - reflects := fun {c} t => by + { reflects := fun {c} t => ⟨by apply IsColimit.ofWhiskerEquivalence e apply isColimitOfReflects F apply IsColimit.ofIsoColimit _ (Functor.mapCoconeWhisker _).symm - exact IsColimit.whiskerEquivalence t _ } + exact IsColimit.whiskerEquivalence t _⟩ } + +@[deprecated "use reflectsColimitsOfShape_of_equiv" (since := "2024-11-19")] +lemma reflectsColimitsOfShapeOfEquiv {J' : Type w₂} [Category.{w₂'} J'] (e : J ≌ J') (F : C ⥤ D) + [ReflectsColimitsOfShape J F] : ReflectsColimitsOfShape J' F := + reflectsColimitsOfShape_of_equiv e F /-- A functor reflecting larger colimits also reflects smaller colimits. -/ -def reflectsColimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] +lemma reflectsColimitsOfSize_of_univLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] [ReflectsColimitsOfSize.{w', w₂'} F] : ReflectsColimitsOfSize.{w, w₂} F where - reflectsColimitsOfShape {J} := reflectsColimitsOfShapeOfEquiv + reflectsColimitsOfShape {J} := reflectsColimitsOfShape_of_equiv ((ShrinkHoms.equivalence J).trans <| Shrink.equivalence _).symm F -/-- `reflectsColimitsOfSizeShrink.{w w'} F` tries to obtain `reflectsColimitsOfSize.{w w'} F` +@[deprecated "use reflectsColimitsOfSize_of_univLE" (since := "2024-11-19")] +lemma reflectsColimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] + [ReflectsColimitsOfSize.{w', w₂'} F] : ReflectsColimitsOfSize.{w, w₂} F := + reflectsColimitsOfSize_of_univLE.{w'} F + +/-- `reflectsColimitsOfSize_shrink.{w w'} F` tries to obtain `reflectsColimitsOfSize.{w w'} F` from some other `reflectsColimitsOfSize F`. -/ -def reflectsColimitsOfSizeShrink (F : C ⥤ D) [ReflectsColimitsOfSize.{max w w₂, max w' w₂'} F] : - ReflectsColimitsOfSize.{w, w'} F := reflectsColimitsOfSizeOfUnivLE.{max w w₂, max w' w₂'} F +lemma reflectsColimitsOfSize_shrink (F : C ⥤ D) [ReflectsColimitsOfSize.{max w w₂, max w' w₂'} F] : + ReflectsColimitsOfSize.{w, w'} F := reflectsColimitsOfSize_of_univLE.{max w w₂, max w' w₂'} F + +@[deprecated "use reflectsColimitsOfSize_shrink" (since := "2024-11-19")] +lemma reflectsColimitsOfSizeShrink (F : C ⥤ D) [ReflectsColimitsOfSize.{max w w₂, max w' w₂'} F] : + ReflectsColimitsOfSize.{w, w'} F := + reflectsColimitsOfSize_shrink F /-- Reflecting colimits at any universe implies reflecting colimits at universe `0`. -/ -def reflectsSmallestColimitsOfReflectsColimits (F : C ⥤ D) [ReflectsColimitsOfSize.{v₃, u₃} F] : +lemma reflectsSmallestColimits_of_reflectsColimits (F : C ⥤ D) [ReflectsColimitsOfSize.{v₃, u₃} F] : ReflectsColimitsOfSize.{0, 0} F := - reflectsColimitsOfSizeShrink F + reflectsColimitsOfSize_shrink F + +@[deprecated "use reflectsSmallestColimits_of_reflectsColimits" (since := "2024-11-19")] +lemma reflectsSmallestColimitsOfReflectsColimits (F : C ⥤ D) [ReflectsColimitsOfSize.{v₃, u₃} F] : + ReflectsColimitsOfSize.{0, 0} F := + reflectsSmallestColimits_of_reflectsColimits F /-- If the colimit of `F` exists and `G` preserves it, then if `G` reflects isomorphisms then it reflects the colimit of `F`. -/ -- Porting note: previous behavior of apply pushed instance holes into hypotheses, this errors -def reflectsColimitOfReflectsIsomorphisms (F : J ⥤ C) (G : C ⥤ D) [G.ReflectsIsomorphisms] +lemma reflectsColimit_of_reflectsIsomorphisms (F : J ⥤ C) (G : C ⥤ D) [G.ReflectsIsomorphisms] [HasColimit F] [PreservesColimit F G] : ReflectsColimit F G where reflects {c} t := by - suffices IsIso (IsColimit.desc (colimit.isColimit F) c) from by - apply IsColimit.ofPointIso (colimit.isColimit F) + suffices IsIso (IsColimit.desc (colimit.isColimit F) c) from ⟨by + apply IsColimit.ofPointIso (colimit.isColimit F)⟩ change IsIso ((Cocones.forget _).map ((colimit.isColimit F).descCoconeMorphism c)) suffices IsIso (IsColimit.descCoconeMorphism (colimit.isColimit F) c) from by apply (Cocones.forget F).map_isIso _ @@ -681,50 +970,74 @@ def reflectsColimitOfReflectsIsomorphisms (F : J ⥤ C) (G : C ⥤ D) [G.Reflect apply isIso_of_reflects_iso _ (Cocones.functoriality F G) exact (isColimitOfPreserves G (colimit.isColimit F)).hom_isIso t _ +@[deprecated "use reflectsColimit_of_reflectsIsomorphisms" (since := "2024-11-19")] +lemma reflectsColimitOfReflectsIsomorphisms (F : J ⥤ C) (G : C ⥤ D) [G.ReflectsIsomorphisms] + [HasColimit F] [PreservesColimit F G] : ReflectsColimit F G := + reflectsColimit_of_reflectsIsomorphisms F G + /-- If `C` has colimits of shape `J` and `G` preserves them, then if `G` reflects isomorphisms then it reflects colimits of shape `J`. -/ -def reflectsColimitsOfShapeOfReflectsIsomorphisms {G : C ⥤ D} [G.ReflectsIsomorphisms] +lemma reflectsColimitsOfShape_of_reflectsIsomorphisms {G : C ⥤ D} [G.ReflectsIsomorphisms] [HasColimitsOfShape J C] [PreservesColimitsOfShape J G] : ReflectsColimitsOfShape J G where - reflectsColimit {F} := reflectsColimitOfReflectsIsomorphisms F G + reflectsColimit {F} := reflectsColimit_of_reflectsIsomorphisms F G + +@[deprecated "use reflectsColimitsOfShape_of_reflectsIsomorphisms" (since := "2024-11-19")] +lemma reflectsColimitsOfShapeOfReflectsIsomorphisms {G : C ⥤ D} [G.ReflectsIsomorphisms] + [HasColimitsOfShape J C] [PreservesColimitsOfShape J G] : ReflectsColimitsOfShape J G := + reflectsColimitsOfShape_of_reflectsIsomorphisms /-- If `C` has colimits and `G` preserves colimits, then if `G` reflects isomorphisms then it reflects colimits. -/ -def reflectsColimitsOfReflectsIsomorphisms {G : C ⥤ D} [G.ReflectsIsomorphisms] +lemma reflectsColimits_of_reflectsIsomorphisms {G : C ⥤ D} [G.ReflectsIsomorphisms] [HasColimitsOfSize.{w', w} C] [PreservesColimitsOfSize.{w', w} G] : ReflectsColimitsOfSize.{w', w} G where - reflectsColimitsOfShape := reflectsColimitsOfShapeOfReflectsIsomorphisms + reflectsColimitsOfShape := reflectsColimitsOfShape_of_reflectsIsomorphisms + +@[deprecated "use reflectsColimits_of_reflectsIsomorphisms" (since := "2024-11-19")] +lemma reflectsColimitsOfReflectsIsomorphisms {G : C ⥤ D} [G.ReflectsIsomorphisms] + [HasColimitsOfSize.{w', w} C] [PreservesColimitsOfSize.{w', w} G] : + ReflectsColimitsOfSize.{w', w} G := + reflectsColimits_of_reflectsIsomorphisms end variable (F : C ⥤ D) /-- A fully faithful functor reflects limits. -/ -instance fullyFaithfulReflectsLimits [F.Full] [F.Faithful] : ReflectsLimitsOfSize.{w, w'} F where +instance fullyFaithful_reflectsLimits [F.Full] [F.Faithful] : ReflectsLimitsOfSize.{w, w'} F where reflectsLimitsOfShape {J} 𝒥₁ := { reflectsLimit := fun {K} => { reflects := fun {c} t => - (IsLimit.mkConeMorphism fun _ => + ⟨(IsLimit.mkConeMorphism fun _ => (Cones.functoriality K F).preimage (t.liftConeMorphism _)) <| by apply fun s m => (Cones.functoriality K F).map_injective _ intro s m rw [Functor.map_preimage] - apply t.uniq_cone_morphism } } + apply t.uniq_cone_morphism⟩ } } + +@[deprecated "use fullyFaithful_reflectsLimits" (since := "2024-11-19")] +lemma fullyFaithfulReflectsLimits [F.Full] [F.Faithful] : ReflectsLimitsOfSize.{w, w'} F := + inferInstance /-- A fully faithful functor reflects colimits. -/ -instance fullyFaithfulReflectsColimits [F.Full] [F.Faithful] : +instance fullyFaithful_reflectsColimits [F.Full] [F.Faithful] : ReflectsColimitsOfSize.{w, w'} F where reflectsColimitsOfShape {J} 𝒥₁ := { reflectsColimit := fun {K} => { reflects := fun {c} t => - (IsColimit.mkCoconeMorphism fun _ => + ⟨(IsColimit.mkCoconeMorphism fun _ => (Cocones.functoriality K F).preimage (t.descCoconeMorphism _)) <| by apply fun s m => (Cocones.functoriality K F).map_injective _ intro s m rw [Functor.map_preimage] - apply t.uniq_cocone_morphism }} + apply t.uniq_cocone_morphism⟩ }} + +@[deprecated "use fullyFaithful_reflectsColimits" (since := "2024-11-19")] +lemma fullyFaithfulReflectsColimits [F.Full] [F.Faithful] : ReflectsColimitsOfSize.{w, w'} F := + inferInstance end CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Filtered.lean b/Mathlib/CategoryTheory/Limits/Preserves/Filtered.lean index d1ad0cedcd3f6..f4b69c1ac9a85 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Filtered.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Filtered.lean @@ -36,7 +36,7 @@ section Preserves /-- `PreservesFilteredColimitsOfSize.{w', w} F` means that `F` sends all colimit cocones over any filtered diagram `J ⥤ C` to colimit cocones, where `J : Type w` with `[Category.{w'} J]`. -/ @[nolint checkUnivs, pp_with_univ] -class PreservesFilteredColimitsOfSize (F : C ⥤ D) where +class PreservesFilteredColimitsOfSize (F : C ⥤ D) : Prop where preserves_filtered_colimits : ∀ (J : Type w) [Category.{w'} J] [IsFiltered J], PreservesColimitsOfShape J F @@ -45,7 +45,7 @@ A functor is said to preserve filtered colimits, if it preserves all colimits of `J` is a filtered category which is small relative to the universe in which morphisms of the source live. -/ -abbrev PreservesFilteredColimits (F : C ⥤ D) : Type max u₁ u₂ (v + 1) := +abbrev PreservesFilteredColimits (F : C ⥤ D) : Prop := PreservesFilteredColimitsOfSize.{v, v} F attribute [instance 100] PreservesFilteredColimitsOfSize.preserves_filtered_colimits @@ -54,35 +54,35 @@ instance (priority := 100) PreservesColimits.preservesFilteredColimits (F : C [PreservesColimitsOfSize.{w, w'} F] : PreservesFilteredColimitsOfSize.{w, w'} F where preserves_filtered_colimits _ := inferInstance -instance compPreservesFilteredColimits (F : C ⥤ D) (G : D ⥤ E) +instance comp_preservesFilteredColimits (F : C ⥤ D) (G : D ⥤ E) [PreservesFilteredColimitsOfSize.{w, w'} F] [PreservesFilteredColimitsOfSize.{w, w'} G] : PreservesFilteredColimitsOfSize.{w, w'} (F ⋙ G) where preserves_filtered_colimits _ := inferInstance /-- A functor preserving larger filtered colimits also preserves smaller filtered colimits. -/ -noncomputable def preservesFilteredColimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] +lemma preservesFilteredColimitsOfSize_of_univLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] [PreservesFilteredColimitsOfSize.{w', w₂'} F] : PreservesFilteredColimitsOfSize.{w, w₂} F where preserves_filtered_colimits J _ _ := by let e := ((ShrinkHoms.equivalence J).trans <| Shrink.equivalence _).symm haveI := IsFiltered.of_equivalence e.symm - exact preservesColimitsOfShapeOfEquiv e F + exact preservesColimitsOfShape_of_equiv e F /-- -`PreservesFilteredColimitsOfSizeShrink.{w, w'} F` tries to obtain +`PreservesFilteredColimitsOfSize_shrink.{w, w'} F` tries to obtain `PreservesFilteredColimitsOfSize.{w, w'} F` from some other `PreservesFilteredColimitsOfSize F`. -/ -noncomputable def preservesFilteredColimitsOfSizeShrink (F : C ⥤ D) +lemma preservesFilteredColimitsOfSize_shrink (F : C ⥤ D) [PreservesFilteredColimitsOfSize.{max w w₂, max w' w₂'} F] : PreservesFilteredColimitsOfSize.{w, w'} F := - preservesFilteredColimitsOfSizeOfUnivLE.{max w w₂, max w' w₂'} F + preservesFilteredColimitsOfSize_of_univLE.{max w w₂, max w' w₂'} F /-- Preserving filtered colimits at any universe implies preserving filtered colimits at universe `0`. -/ -noncomputable def preservesSmallestFilteredColimitsOfPreservesFilteredColimits (F : C ⥤ D) +lemma preservesSmallestFilteredColimits_of_preservesFilteredColimits (F : C ⥤ D) [PreservesFilteredColimitsOfSize.{w', w} F] : PreservesFilteredColimitsOfSize.{0, 0} F := - preservesFilteredColimitsOfSizeShrink F + preservesFilteredColimitsOfSize_shrink F end Preserves @@ -92,7 +92,7 @@ section Reflects /-- `ReflectsFilteredColimitsOfSize.{w', w} F` means that whenever the image of a filtered cocone under `F` is a colimit cocone, the original cocone was already a colimit. -/ @[nolint checkUnivs, pp_with_univ] -class ReflectsFilteredColimitsOfSize (F : C ⥤ D) where +class ReflectsFilteredColimitsOfSize (F : C ⥤ D) : Prop where reflects_filtered_colimits : ∀ (J : Type w) [Category.{w'} J] [IsFiltered J], ReflectsColimitsOfShape J F @@ -101,7 +101,7 @@ A functor is said to reflect filtered colimits, if it reflects all colimits of s `J` is a filtered category which is small relative to the universe in which morphisms of the source live. -/ -abbrev ReflectsFilteredColimits (F : C ⥤ D) : Type max u₁ u₂ (v + 1) := +abbrev ReflectsFilteredColimits (F : C ⥤ D) : Prop := ReflectsFilteredColimitsOfSize.{v, v} F attribute [instance 100] ReflectsFilteredColimitsOfSize.reflects_filtered_colimits @@ -110,35 +110,35 @@ instance (priority := 100) ReflectsColimits.reflectsFilteredColimits (F : C ⥤ [ReflectsColimitsOfSize.{w, w'} F] : ReflectsFilteredColimitsOfSize.{w, w'} F where reflects_filtered_colimits _ := inferInstance -instance compReflectsFilteredColimits (F : C ⥤ D) (G : D ⥤ E) +instance comp_reflectsFilteredColimits (F : C ⥤ D) (G : D ⥤ E) [ReflectsFilteredColimitsOfSize.{w, w'} F] [ReflectsFilteredColimitsOfSize.{w, w'} G] : ReflectsFilteredColimitsOfSize.{w, w'} (F ⋙ G) where reflects_filtered_colimits _ := inferInstance /-- A functor reflecting larger filtered colimits also reflects smaller filtered colimits. -/ -noncomputable def reflectsFilteredColimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] +lemma reflectsFilteredColimitsOfSize_of_univLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] [ReflectsFilteredColimitsOfSize.{w', w₂'} F] : ReflectsFilteredColimitsOfSize.{w, w₂} F where reflects_filtered_colimits J _ _ := by let e := ((ShrinkHoms.equivalence J).trans <| Shrink.equivalence _).symm haveI := IsFiltered.of_equivalence e.symm - exact reflectsColimitsOfShapeOfEquiv e F + exact reflectsColimitsOfShape_of_equiv e F /-- -`ReflectsFilteredColimitsOfSizeShrink.{w, w'} F` tries to obtain +`ReflectsFilteredColimitsOfSize_shrink.{w, w'} F` tries to obtain `ReflectsFilteredColimitsOfSize.{w, w'} F` from some other `ReflectsFilteredColimitsOfSize F`. -/ -noncomputable def reflectsFilteredColimitsOfSizeShrink (F : C ⥤ D) +lemma reflectsFilteredColimitsOfSize_shrink (F : C ⥤ D) [ReflectsFilteredColimitsOfSize.{max w w₂, max w' w₂'} F] : ReflectsFilteredColimitsOfSize.{w, w'} F := - reflectsFilteredColimitsOfSizeOfUnivLE.{max w w₂, max w' w₂'} F + reflectsFilteredColimitsOfSize_of_univLE.{max w w₂, max w' w₂'} F /-- Reflecting filtered colimits at any universe implies reflecting filtered colimits at universe `0`. -/ -noncomputable def reflectsSmallestFilteredColimitsOfReflectsFilteredColimits (F : C ⥤ D) +lemma reflectsSmallestFilteredColimits_of_reflectsFilteredColimits (F : C ⥤ D) [ReflectsFilteredColimitsOfSize.{w', w} F] : ReflectsFilteredColimitsOfSize.{0, 0} F := - reflectsFilteredColimitsOfSizeShrink F + reflectsFilteredColimitsOfSize_shrink F end Reflects @@ -152,7 +152,7 @@ section Preserves /-- `PreservesCofilteredLimitsOfSize.{w', w} F` means that `F` sends all limit cones over any cofiltered diagram `J ⥤ C` to limit cones, where `J : Type w` with `[Category.{w'} J]`. -/ @[nolint checkUnivs, pp_with_univ] -class PreservesCofilteredLimitsOfSize (F : C ⥤ D) where +class PreservesCofilteredLimitsOfSize (F : C ⥤ D) : Prop where preserves_cofiltered_limits : ∀ (J : Type w) [Category.{w'} J] [IsCofiltered J], PreservesLimitsOfShape J F @@ -161,7 +161,7 @@ A functor is said to preserve cofiltered limits, if it preserves all limits of s `J` is a cofiltered category which is small relative to the universe in which morphisms of the source live. -/ -abbrev PreservesCofilteredLimits (F : C ⥤ D) : Type max u₁ u₂ (v + 1) := +abbrev PreservesCofilteredLimits (F : C ⥤ D) : Prop := PreservesCofilteredLimitsOfSize.{v, v} F attribute [instance 100] PreservesCofilteredLimitsOfSize.preserves_cofiltered_limits @@ -170,35 +170,35 @@ instance (priority := 100) PreservesLimits.preservesCofilteredLimits (F : C ⥤ [PreservesLimitsOfSize.{w, w'} F] : PreservesCofilteredLimitsOfSize.{w, w'} F where preserves_cofiltered_limits _ := inferInstance -instance compPreservesCofilteredLimits (F : C ⥤ D) (G : D ⥤ E) +instance comp_preservesCofilteredLimits (F : C ⥤ D) (G : D ⥤ E) [PreservesCofilteredLimitsOfSize.{w, w'} F] [PreservesCofilteredLimitsOfSize.{w, w'} G] : PreservesCofilteredLimitsOfSize.{w, w'} (F ⋙ G) where preserves_cofiltered_limits _ := inferInstance /-- A functor preserving larger cofiltered limits also preserves smaller cofiltered limits. -/ -noncomputable def preservesCofilteredLimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] +lemma preservesCofilteredLimitsOfSize_of_univLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] [PreservesCofilteredLimitsOfSize.{w', w₂'} F] : PreservesCofilteredLimitsOfSize.{w, w₂} F where preserves_cofiltered_limits J _ _ := by let e := ((ShrinkHoms.equivalence J).trans <| Shrink.equivalence _).symm haveI := IsCofiltered.of_equivalence e.symm - exact preservesLimitsOfShapeOfEquiv e F + exact preservesLimitsOfShape_of_equiv e F /-- `PreservesCofilteredLimitsOfSizeShrink.{w, w'} F` tries to obtain `PreservesCofilteredLimitsOfSize.{w, w'} F` from some other `PreservesCofilteredLimitsOfSize F`. -/ -noncomputable def preservesCofilteredLimitsOfSizeShrink (F : C ⥤ D) +lemma preservesCofilteredLimitsOfSize_shrink (F : C ⥤ D) [PreservesCofilteredLimitsOfSize.{max w w₂, max w' w₂'} F] : PreservesCofilteredLimitsOfSize.{w, w'} F := - preservesCofilteredLimitsOfSizeOfUnivLE.{max w w₂, max w' w₂'} F + preservesCofilteredLimitsOfSize_of_univLE.{max w w₂, max w' w₂'} F /-- Preserving cofiltered limits at any universe implies preserving cofiltered limits at universe `0`. -/ -noncomputable def preservesSmallestCofilteredLimitsOfPreservesCofilteredLimits (F : C ⥤ D) +lemma preservesSmallestCofilteredLimits_of_preservesCofilteredLimits (F : C ⥤ D) [PreservesCofilteredLimitsOfSize.{w', w} F] : PreservesCofilteredLimitsOfSize.{0, 0} F := - preservesCofilteredLimitsOfSizeShrink F + preservesCofilteredLimitsOfSize_shrink F end Preserves @@ -208,7 +208,7 @@ section Reflects /-- `ReflectsCofilteredLimitsOfSize.{w', w} F` means that whenever the image of a cofiltered cone under `F` is a limit cone, the original cone was already a limit. -/ @[nolint checkUnivs, pp_with_univ] -class ReflectsCofilteredLimitsOfSize (F : C ⥤ D) where +class ReflectsCofilteredLimitsOfSize (F : C ⥤ D) : Prop where reflects_cofiltered_limits : ∀ (J : Type w) [Category.{w'} J] [IsCofiltered J], ReflectsLimitsOfShape J F @@ -217,7 +217,7 @@ A functor is said to reflect cofiltered limits, if it reflects all limits of sha `J` is a cofiltered category which is small relative to the universe in which morphisms of the source live. -/ -abbrev ReflectsCofilteredLimits (F : C ⥤ D) : Type max u₁ u₂ (v + 1) := +abbrev ReflectsCofilteredLimits (F : C ⥤ D) : Prop := ReflectsCofilteredLimitsOfSize.{v, v} F attribute [instance 100] ReflectsCofilteredLimitsOfSize.reflects_cofiltered_limits @@ -226,35 +226,35 @@ instance (priority := 100) ReflectsLimits.reflectsCofilteredLimits (F : C ⥤ D) [ReflectsLimitsOfSize.{w, w'} F] : ReflectsCofilteredLimitsOfSize.{w, w'} F where reflects_cofiltered_limits _ := inferInstance -instance compReflectsCofilteredLimits (F : C ⥤ D) (G : D ⥤ E) +instance comp_reflectsCofilteredLimits (F : C ⥤ D) (G : D ⥤ E) [ReflectsCofilteredLimitsOfSize.{w, w'} F] [ReflectsCofilteredLimitsOfSize.{w, w'} G] : ReflectsCofilteredLimitsOfSize.{w, w'} (F ⋙ G) where reflects_cofiltered_limits _ := inferInstance /-- A functor reflecting larger cofiltered limits also reflects smaller cofiltered limits. -/ -noncomputable def reflectsCofilteredLimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] +lemma reflectsCofilteredLimitsOfSize_of_univLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] [ReflectsCofilteredLimitsOfSize.{w', w₂'} F] : ReflectsCofilteredLimitsOfSize.{w, w₂} F where reflects_cofiltered_limits J _ _ := by let e := ((ShrinkHoms.equivalence J).trans <| Shrink.equivalence _).symm haveI := IsCofiltered.of_equivalence e.symm - exact reflectsLimitsOfShapeOfEquiv e F + exact reflectsLimitsOfShape_of_equiv e F /-- -`ReflectsCofilteredLimitsOfSizeShrink.{w, w'} F` tries to obtain +`ReflectsCofilteredLimitsOfSize_shrink.{w, w'} F` tries to obtain `ReflectsCofilteredLimitsOfSize.{w, w'} F` from some other `ReflectsCofilteredLimitsOfSize F`. -/ -noncomputable def reflectsCofilteredLimitsOfSizeShrink (F : C ⥤ D) +lemma reflectsCofilteredLimitsOfSize_shrink (F : C ⥤ D) [ReflectsCofilteredLimitsOfSize.{max w w₂, max w' w₂'} F] : ReflectsCofilteredLimitsOfSize.{w, w'} F := - reflectsCofilteredLimitsOfSizeOfUnivLE.{max w w₂, max w' w₂'} F + reflectsCofilteredLimitsOfSize_of_univLE.{max w w₂, max w' w₂'} F /-- Reflecting cofiltered limits at any universe implies reflecting cofiltered limits at universe `0`. -/ -noncomputable def reflectsSmallestCofilteredLimitsOfReflectsCofilteredLimits (F : C ⥤ D) +lemma reflectsSmallestCofilteredLimits_of_reflectsCofilteredLimits (F : C ⥤ D) [ReflectsCofilteredLimitsOfSize.{w', w} F] : ReflectsCofilteredLimitsOfSize.{0, 0} F := - reflectsCofilteredLimitsOfSizeShrink F + reflectsCofilteredLimitsOfSize_shrink F end Reflects diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Finite.lean b/Mathlib/CategoryTheory/Limits/Preserves/Finite.lean index e852cc4007abb..89f2b68d3e856 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Finite.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Finite.lean @@ -37,127 +37,127 @@ variable {J : Type w} [SmallCategory J] {K : J ⥤ C} /-- A functor is said to preserve finite limits, if it preserves all limits of shape `J`, where `J : Type` is a finite category. -/ -class PreservesFiniteLimits (F : C ⥤ D) where +class PreservesFiniteLimits (F : C ⥤ D) : Prop where preservesFiniteLimits : ∀ (J : Type) [SmallCategory J] [FinCategory J], PreservesLimitsOfShape J F := by infer_instance attribute [instance] PreservesFiniteLimits.preservesFiniteLimits instance (F : C ⥤ D) : Subsingleton (PreservesFiniteLimits F) := - ⟨fun ⟨a⟩ ⟨b⟩ => by congr; subsingleton⟩ + ⟨fun ⟨a⟩ ⟨b⟩ => by congr⟩ /-- Preserving finite limits also implies preserving limits over finite shapes in higher universes, though through a noncomputable instance. -/ -noncomputable instance (priority := 100) preservesLimitsOfShapeOfPreservesFiniteLimits (F : C ⥤ D) +instance (priority := 100) preservesLimitsOfShapeOfPreservesFiniteLimits (F : C ⥤ D) [PreservesFiniteLimits F] (J : Type w) [SmallCategory J] [FinCategory J] : PreservesLimitsOfShape J F := by - apply preservesLimitsOfShapeOfEquiv (FinCategory.equivAsType J) + apply preservesLimitsOfShape_of_equiv (FinCategory.equivAsType J) -- This is a dangerous instance as it has unbound universe variables. /-- If we preserve limits of some arbitrary size, then we preserve all finite limits. -/ -noncomputable def PreservesLimitsOfSize.preservesFiniteLimits (F : C ⥤ D) +lemma PreservesLimitsOfSize.preservesFiniteLimits (F : C ⥤ D) [PreservesLimitsOfSize.{w, w₂} F] : PreservesFiniteLimits F where preservesFiniteLimits J (sJ : SmallCategory J) fJ := by - haveI := preservesSmallestLimitsOfPreservesLimits F - exact preservesLimitsOfShapeOfEquiv (FinCategory.equivAsType J) F + haveI := preservesSmallestLimits_of_preservesLimits F + exact preservesLimitsOfShape_of_equiv (FinCategory.equivAsType J) F -- Added as a specialization of the dangerous instance above, for limits indexed in Type 0. -noncomputable instance (priority := 120) PreservesLimitsOfSize0.preservesFiniteLimits +instance (priority := 120) PreservesLimitsOfSize0.preservesFiniteLimits (F : C ⥤ D) [PreservesLimitsOfSize.{0, 0} F] : PreservesFiniteLimits F := PreservesLimitsOfSize.preservesFiniteLimits F -- An alternative specialization of the dangerous instance for small limits. -noncomputable instance (priority := 120) PreservesLimits.preservesFiniteLimits (F : C ⥤ D) +instance (priority := 120) PreservesLimits.preservesFiniteLimits (F : C ⥤ D) [PreservesLimits F] : PreservesFiniteLimits F := PreservesLimitsOfSize.preservesFiniteLimits F /-- We can always derive `PreservesFiniteLimits C` by showing that we are preserving limits at an arbitrary universe. -/ -def preservesFiniteLimitsOfPreservesFiniteLimitsOfSize (F : C ⥤ D) +lemma preservesFiniteLimits_of_preservesFiniteLimitsOfSize (F : C ⥤ D) (h : ∀ (J : Type w) {𝒥 : SmallCategory J} (_ : @FinCategory J 𝒥), PreservesLimitsOfShape J F) : PreservesFiniteLimits F where preservesFiniteLimits J (_ : SmallCategory J) _ := by letI : Category (ULiftHom (ULift J)) := ULiftHom.category haveI := h (ULiftHom (ULift J)) CategoryTheory.finCategoryUlift - exact preservesLimitsOfShapeOfEquiv (ULiftHomULiftCategory.equiv J).symm F + exact preservesLimitsOfShape_of_equiv (ULiftHomULiftCategory.equiv J).symm F /-- The composition of two left exact functors is left exact. -/ -def compPreservesFiniteLimits (F : C ⥤ D) (G : D ⥤ E) [PreservesFiniteLimits F] +lemma comp_preservesFiniteLimits (F : C ⥤ D) (G : D ⥤ E) [PreservesFiniteLimits F] [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] : +lemma preservesFiniteLimits_of_natIso {F G : C ⥤ D} (h : F ≅ G) [PreservesFiniteLimits F] : PreservesFiniteLimits G where - preservesFiniteLimits _ _ _ := preservesLimitsOfShapeOfNatIso h + preservesFiniteLimits _ _ _ := preservesLimitsOfShape_of_natIso h /- Porting note: adding this class because quantified classes don't behave well https://github.com/leanprover-community/mathlib4/pull/2764 -/ /-- A functor `F` preserves finite products if it preserves all from `Discrete J` for `Fintype J` -/ -class PreservesFiniteProducts (F : C ⥤ D) where +class PreservesFiniteProducts (F : C ⥤ D) : Prop where preserves : ∀ (J : Type) [Fintype J], PreservesLimitsOfShape (Discrete J) F attribute [instance] PreservesFiniteProducts.preserves instance (F : C ⥤ D) : Subsingleton (PreservesFiniteProducts F) := - ⟨fun ⟨a⟩ ⟨b⟩ => by congr; subsingleton⟩ + ⟨fun ⟨a⟩ ⟨b⟩ => by congr⟩ -noncomputable instance (priority := 100) (F : C ⥤ D) (J : Type u) [Finite J] +instance (priority := 100) (F : C ⥤ D) (J : Type u) [Finite J] [PreservesFiniteProducts F] : PreservesLimitsOfShape (Discrete J) F := by apply Nonempty.some obtain ⟨n, ⟨e⟩⟩ := Finite.exists_equiv_fin J - exact ⟨preservesLimitsOfShapeOfEquiv (Discrete.equivalence e.symm) F⟩ + exact ⟨preservesLimitsOfShape_of_equiv (Discrete.equivalence e.symm) F⟩ -instance compPreservesFiniteProducts (F : C ⥤ D) (G : D ⥤ E) +instance comp_preservesFiniteProducts (F : C ⥤ D) (G : D ⥤ E) [PreservesFiniteProducts F] [PreservesFiniteProducts G] : PreservesFiniteProducts (F ⋙ G) where preserves _ _ := inferInstance -noncomputable instance (F : C ⥤ D) [PreservesFiniteLimits F] : PreservesFiniteProducts F where +instance (F : C ⥤ D) [PreservesFiniteLimits F] : PreservesFiniteProducts F where preserves _ _ := inferInstance /-- A functor is said to reflect finite limits, if it reflects all limits of shape `J`, where `J : Type` is a finite category. -/ -class ReflectsFiniteLimits (F : C ⥤ D) where +class ReflectsFiniteLimits (F : C ⥤ D) : Prop where reflects : ∀ (J : Type) [SmallCategory J] [FinCategory J], ReflectsLimitsOfShape J F := by infer_instance attribute [instance] ReflectsFiniteLimits.reflects instance (F : C ⥤ D) : Subsingleton (ReflectsFiniteLimits F) := - ⟨fun ⟨a⟩ ⟨b⟩ => by congr; subsingleton⟩ + ⟨fun ⟨a⟩ ⟨b⟩ => by congr⟩ /- Similarly to preserving finite products, quantified classes don't behave well. -/ /-- A functor `F` preserves finite products if it reflects limits of shape `Discrete J` for finite `J` -/ -class ReflectsFiniteProducts (F : C ⥤ D) where +class ReflectsFiniteProducts (F : C ⥤ D) : Prop where reflects : ∀ (J : Type) [Fintype J], ReflectsLimitsOfShape (Discrete J) F attribute [instance] ReflectsFiniteProducts.reflects instance (F : C ⥤ D) : Subsingleton (ReflectsFiniteProducts F) := - ⟨fun ⟨a⟩ ⟨b⟩ => by congr; subsingleton⟩ + ⟨fun ⟨a⟩ ⟨b⟩ => by congr⟩ -- This is a dangerous instance as it has unbound universe variables. /-- If we reflect limits of some arbitrary size, then we reflect all finite limits. -/ -noncomputable def ReflectsLimitsOfSize.reflectsFiniteLimits +lemma ReflectsLimitsOfSize.reflectsFiniteLimits (F : C ⥤ D) [ReflectsLimitsOfSize.{w, w₂} F] : ReflectsFiniteLimits F where reflects J (sJ : SmallCategory J) fJ := by - haveI := reflectsSmallestLimitsOfReflectsLimits F - exact reflectsLimitsOfShapeOfEquiv (FinCategory.equivAsType J) F + haveI := reflectsSmallestLimits_of_reflectsLimits F + exact reflectsLimitsOfShape_of_equiv (FinCategory.equivAsType J) F -- Added as a specialization of the dangerous instance above, for colimits indexed in Type 0. -noncomputable instance (priority := 120) (F : C ⥤ D) [ReflectsLimitsOfSize.{0, 0} F] : +instance (priority := 120) (F : C ⥤ D) [ReflectsLimitsOfSize.{0, 0} F] : ReflectsFiniteLimits F := ReflectsLimitsOfSize.reflectsFiniteLimits F -- An alternative specialization of the dangerous instance for small colimits. -noncomputable instance (priority := 120) (F : C ⥤ D) +instance (priority := 120) (F : C ⥤ D) [ReflectsLimits F] : ReflectsFiniteLimits F := ReflectsLimitsOfSize.reflectsFiniteLimits F @@ -165,40 +165,40 @@ noncomputable instance (priority := 120) (F : C ⥤ D) If `F ⋙ G` preserves finite limits and `G` reflects finite limits, then `F` preserves finite limits. -/ -def preservesFiniteLimitsOfReflectsOfPreserves (F : C ⥤ D) (G : D ⥤ E) +lemma preservesFiniteLimits_of_reflects_of_preserves (F : C ⥤ D) (G : D ⥤ E) [PreservesFiniteLimits (F ⋙ G)] [ReflectsFiniteLimits G] : PreservesFiniteLimits F where - preservesFiniteLimits _ _ _ := preservesLimitsOfShapeOfReflectsOfPreserves F G + preservesFiniteLimits _ _ _ := preservesLimitsOfShape_of_reflects_of_preserves F G /-- If `F ⋙ G` preserves finite products and `G` reflects finite products, then `F` preserves finite products. -/ -def preservesFiniteProductsOfReflectsOfPreserves (F : C ⥤ D) (G : D ⥤ E) +lemma preservesFiniteProducts_of_reflects_of_preserves (F : C ⥤ D) (G : D ⥤ E) [PreservesFiniteProducts (F ⋙ G)] [ReflectsFiniteProducts G] : PreservesFiniteProducts F where - preserves _ _ := preservesLimitsOfShapeOfReflectsOfPreserves F G + preserves _ _ := preservesLimitsOfShape_of_reflects_of_preserves F G -noncomputable instance reflectsFiniteLimitsOfReflectsIsomorphisms (F : C ⥤ D) +instance reflectsFiniteLimits_of_reflectsIsomorphisms (F : C ⥤ D) [F.ReflectsIsomorphisms] [HasFiniteLimits C] [PreservesFiniteLimits F] : ReflectsFiniteLimits F where - reflects _ _ _ := reflectsLimitsOfShapeOfReflectsIsomorphisms + reflects _ _ _ := reflectsLimitsOfShape_of_reflectsIsomorphisms -noncomputable instance reflectsFiniteProductsOfReflectsIsomorphisms (F : C ⥤ D) +instance reflectsFiniteProducts_of_reflectsIsomorphisms (F : C ⥤ D) [F.ReflectsIsomorphisms] [HasFiniteProducts C] [PreservesFiniteProducts F] : ReflectsFiniteProducts F where - reflects _ _ := reflectsLimitsOfShapeOfReflectsIsomorphisms + reflects _ _ := reflectsLimitsOfShape_of_reflectsIsomorphisms -instance compReflectsFiniteProducts (F : C ⥤ D) (G : D ⥤ E) +instance comp_reflectsFiniteProducts (F : C ⥤ D) (G : D ⥤ E) [ReflectsFiniteProducts F] [ReflectsFiniteProducts G] : ReflectsFiniteProducts (F ⋙ G) where reflects _ _ := inferInstance -noncomputable instance (F : C ⥤ D) [ReflectsFiniteLimits F] : ReflectsFiniteProducts F where +instance (F : C ⥤ D) [ReflectsFiniteLimits F] : ReflectsFiniteProducts F where reflects _ _ := inferInstance /-- A functor is said to preserve finite colimits, if it preserves all colimits of shape `J`, where `J : Type` is a finite category. -/ -class PreservesFiniteColimits (F : C ⥤ D) where +class PreservesFiniteColimits (F : C ⥤ D) : Prop where preservesFiniteColimits : ∀ (J : Type) [SmallCategory J] [FinCategory J], PreservesColimitsOfShape J F := by infer_instance @@ -206,111 +206,111 @@ class PreservesFiniteColimits (F : C ⥤ D) where attribute [instance] PreservesFiniteColimits.preservesFiniteColimits instance (F : C ⥤ D) : Subsingleton (PreservesFiniteColimits F) := - ⟨fun ⟨a⟩ ⟨b⟩ => by congr; subsingleton⟩ + ⟨fun ⟨a⟩ ⟨b⟩ => by congr⟩ /-- Preserving finite colimits also implies preserving colimits over finite shapes in higher -universes, though through a noncomputable instance. +universes. -/ -noncomputable instance (priority := 100) preservesColimitsOfShapeOfPreservesFiniteColimits +instance (priority := 100) preservesColimitsOfShapeOfPreservesFiniteColimits (F : C ⥤ D) [PreservesFiniteColimits F] (J : Type w) [SmallCategory J] [FinCategory J] : PreservesColimitsOfShape J F := by - apply preservesColimitsOfShapeOfEquiv (FinCategory.equivAsType J) + apply preservesColimitsOfShape_of_equiv (FinCategory.equivAsType J) -- This is a dangerous instance as it has unbound universe variables. /-- If we preserve colimits of some arbitrary size, then we preserve all finite colimits. -/ -noncomputable def PreservesColimitsOfSize.preservesFiniteColimits (F : C ⥤ D) +lemma PreservesColimitsOfSize.preservesFiniteColimits (F : C ⥤ D) [PreservesColimitsOfSize.{w, w₂} F] : PreservesFiniteColimits F where preservesFiniteColimits J (sJ : SmallCategory J) fJ := by - haveI := preservesSmallestColimitsOfPreservesColimits F - exact preservesColimitsOfShapeOfEquiv (FinCategory.equivAsType J) F + haveI := preservesSmallestColimits_of_preservesColimits F + exact preservesColimitsOfShape_of_equiv (FinCategory.equivAsType J) F -- Added as a specialization of the dangerous instance above, for colimits indexed in Type 0. -noncomputable instance (priority := 120) PreservesColimitsOfSize0.preservesFiniteColimits +instance (priority := 120) PreservesColimitsOfSize0.preservesFiniteColimits (F : C ⥤ D) [PreservesColimitsOfSize.{0, 0} F] : PreservesFiniteColimits F := PreservesColimitsOfSize.preservesFiniteColimits F -- An alternative specialization of the dangerous instance for small colimits. -noncomputable instance (priority := 120) PreservesColimits.preservesFiniteColimits (F : C ⥤ D) +instance (priority := 120) PreservesColimits.preservesFiniteColimits (F : C ⥤ D) [PreservesColimits F] : PreservesFiniteColimits F := PreservesColimitsOfSize.preservesFiniteColimits F /-- We can always derive `PreservesFiniteColimits C` by showing that we are preserving colimits at an arbitrary universe. -/ -def preservesFiniteColimitsOfPreservesFiniteColimitsOfSize (F : C ⥤ D) +lemma preservesFiniteColimits_of_preservesFiniteColimitsOfSize (F : C ⥤ D) (h : ∀ (J : Type w) {𝒥 : SmallCategory J} (_ : @FinCategory J 𝒥), PreservesColimitsOfShape J F) : PreservesFiniteColimits F where preservesFiniteColimits J (_ : SmallCategory J) _ := by letI : Category (ULiftHom (ULift J)) := ULiftHom.category haveI := h (ULiftHom (ULift J)) CategoryTheory.finCategoryUlift - exact preservesColimitsOfShapeOfEquiv (ULiftHomULiftCategory.equiv J).symm F + exact preservesColimitsOfShape_of_equiv (ULiftHomULiftCategory.equiv J).symm F /-- The composition of two right exact functors is right exact. -/ -def compPreservesFiniteColimits (F : C ⥤ D) (G : D ⥤ E) [PreservesFiniteColimits F] +lemma comp_preservesFiniteColimits (F : C ⥤ D) (G : D ⥤ E) [PreservesFiniteColimits F] [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] : +lemma preservesFiniteColimits_of_natIso {F G : C ⥤ D} (h : F ≅ G) [PreservesFiniteColimits F] : PreservesFiniteColimits G where - preservesFiniteColimits _ _ _ := preservesColimitsOfShapeOfNatIso h + preservesFiniteColimits _ _ _ := preservesColimitsOfShape_of_natIso h /- Porting note: adding this class because quantified classes don't behave well https://github.com/leanprover-community/mathlib4/pull/2764 -/ /-- A functor `F` preserves finite products if it preserves all from `Discrete J` for `Fintype J` -/ -class PreservesFiniteCoproducts (F : C ⥤ D) where +class PreservesFiniteCoproducts (F : C ⥤ D) : Prop where /-- preservation of colimits indexed by `Discrete J` when `[Fintype J]` -/ preserves : ∀ (J : Type) [Fintype J], PreservesColimitsOfShape (Discrete J) F attribute [instance] PreservesFiniteCoproducts.preserves instance (F : C ⥤ D) : Subsingleton (PreservesFiniteCoproducts F) := - ⟨fun ⟨a⟩ ⟨b⟩ => by congr; subsingleton⟩ + ⟨fun ⟨a⟩ ⟨b⟩ => by congr⟩ -noncomputable instance (priority := 100) (F : C ⥤ D) (J : Type u) [Finite J] +instance (priority := 100) (F : C ⥤ D) (J : Type u) [Finite J] [PreservesFiniteCoproducts F] : PreservesColimitsOfShape (Discrete J) F := by apply Nonempty.some obtain ⟨n, ⟨e⟩⟩ := Finite.exists_equiv_fin J - exact ⟨preservesColimitsOfShapeOfEquiv (Discrete.equivalence e.symm) F⟩ + exact ⟨preservesColimitsOfShape_of_equiv (Discrete.equivalence e.symm) F⟩ -instance compPreservesFiniteCoproducts (F : C ⥤ D) (G : D ⥤ E) +instance comp_preservesFiniteCoproducts (F : C ⥤ D) (G : D ⥤ E) [PreservesFiniteCoproducts F] [PreservesFiniteCoproducts G] : PreservesFiniteCoproducts (F ⋙ G) where preserves _ _ := inferInstance -noncomputable instance (F : C ⥤ D) [PreservesFiniteColimits F] : PreservesFiniteCoproducts F where +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. -/ -class ReflectsFiniteColimits (F : C ⥤ D) where +class ReflectsFiniteColimits (F : C ⥤ D) : Prop where reflects : ∀ (J : Type) [SmallCategory J] [FinCategory J], ReflectsColimitsOfShape J F := by infer_instance attribute [instance] ReflectsFiniteColimits.reflects instance (F : C ⥤ D) : Subsingleton (ReflectsFiniteColimits F) := - ⟨fun ⟨a⟩ ⟨b⟩ => by congr; subsingleton⟩ + ⟨fun ⟨a⟩ ⟨b⟩ => by congr⟩ -- This is a dangerous instance as it has unbound universe variables. /-- If we reflect colimits of some arbitrary size, then we reflect all finite colimits. -/ -noncomputable def ReflectsColimitsOfSize.reflectsFiniteColimits +lemma ReflectsColimitsOfSize.reflectsFiniteColimits (F : C ⥤ D) [ReflectsColimitsOfSize.{w, w₂} F] : ReflectsFiniteColimits F where reflects J (sJ : SmallCategory J) fJ := by - haveI := reflectsSmallestColimitsOfReflectsColimits F - exact reflectsColimitsOfShapeOfEquiv (FinCategory.equivAsType J) F + haveI := reflectsSmallestColimits_of_reflectsColimits F + exact reflectsColimitsOfShape_of_equiv (FinCategory.equivAsType J) F -- Added as a specialization of the dangerous instance above, for colimits indexed in Type 0. -noncomputable instance (priority := 120) (F : C ⥤ D) [ReflectsColimitsOfSize.{0, 0} F] : +instance (priority := 120) (F : C ⥤ D) [ReflectsColimitsOfSize.{0, 0} F] : ReflectsFiniteColimits F := ReflectsColimitsOfSize.reflectsFiniteColimits F -- An alternative specialization of the dangerous instance for small colimits. -noncomputable instance (priority := 120) (F : C ⥤ D) +instance (priority := 120) (F : C ⥤ D) [ReflectsColimits F] : ReflectsFiniteColimits F := ReflectsColimitsOfSize.reflectsFiniteColimits F @@ -319,47 +319,47 @@ noncomputable instance (priority := 120) (F : C ⥤ D) A functor `F` preserves finite coproducts if it reflects colimits of shape `Discrete J` for finite `J` -/ -class ReflectsFiniteCoproducts (F : C ⥤ D) where +class ReflectsFiniteCoproducts (F : C ⥤ D) : Prop where reflects : ∀ (J : Type) [Fintype J], ReflectsColimitsOfShape (Discrete J) F attribute [instance] ReflectsFiniteCoproducts.reflects instance (F : C ⥤ D) : Subsingleton (ReflectsFiniteCoproducts F) := - ⟨fun ⟨a⟩ ⟨b⟩ => by congr; subsingleton⟩ + ⟨fun ⟨a⟩ ⟨b⟩ => by congr⟩ /-- If `F ⋙ G` preserves finite colimits and `G` reflects finite colimits, then `F` preserves finite colimits. -/ -def preservesFiniteColimitsOfReflectsOfPreserves (F : C ⥤ D) (G : D ⥤ E) +lemma preservesFiniteColimits_of_reflects_of_preserves (F : C ⥤ D) (G : D ⥤ E) [PreservesFiniteColimits (F ⋙ G)] [ReflectsFiniteColimits G] : PreservesFiniteColimits F where - preservesFiniteColimits _ _ _ := preservesColimitsOfShapeOfReflectsOfPreserves F G + preservesFiniteColimits _ _ _ := preservesColimitsOfShape_of_reflects_of_preserves F G /-- If `F ⋙ G` preserves finite coproducts and `G` reflects finite coproducts, then `F` preserves finite coproducts. -/ -noncomputable def preservesFiniteCoproductsOfReflectsOfPreserves (F : C ⥤ D) (G : D ⥤ E) +lemma preservesFiniteCoproducts_of_reflects_of_preserves (F : C ⥤ D) (G : D ⥤ E) [PreservesFiniteCoproducts (F ⋙ G)] [ReflectsFiniteCoproducts G] : PreservesFiniteCoproducts F where - preserves _ _ := preservesColimitsOfShapeOfReflectsOfPreserves F G + preserves _ _ := preservesColimitsOfShape_of_reflects_of_preserves F G -noncomputable instance reflectsFiniteColimitsOfReflectsIsomorphisms (F : C ⥤ D) +instance reflectsFiniteColimitsOfReflectsIsomorphisms (F : C ⥤ D) [F.ReflectsIsomorphisms] [HasFiniteColimits C] [PreservesFiniteColimits F] : ReflectsFiniteColimits F where - reflects _ _ _ := reflectsColimitsOfShapeOfReflectsIsomorphisms + reflects _ _ _ := reflectsColimitsOfShape_of_reflectsIsomorphisms -noncomputable instance reflectsFiniteCoproductsOfReflectsIsomorphisms (F : C ⥤ D) +instance reflectsFiniteCoproductsOfReflectsIsomorphisms (F : C ⥤ D) [F.ReflectsIsomorphisms] [HasFiniteCoproducts C] [PreservesFiniteCoproducts F] : ReflectsFiniteCoproducts F where - reflects _ _ := reflectsColimitsOfShapeOfReflectsIsomorphisms + reflects _ _ := reflectsColimitsOfShape_of_reflectsIsomorphisms -instance compReflectsFiniteCoproducts (F : C ⥤ D) (G : D ⥤ E) +instance comp_reflectsFiniteCoproducts (F : C ⥤ D) (G : D ⥤ E) [ReflectsFiniteCoproducts F] [ReflectsFiniteCoproducts G] : ReflectsFiniteCoproducts (F ⋙ G) where reflects _ _ := inferInstance -noncomputable instance (F : C ⥤ D) [ReflectsFiniteColimits F] : ReflectsFiniteCoproducts F where +instance (F : C ⥤ D) [ReflectsFiniteColimits F] : ReflectsFiniteCoproducts F where reflects _ _ := inferInstance end CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Limits/Preserves/FunctorCategory.lean b/Mathlib/CategoryTheory/Limits/Preserves/FunctorCategory.lean index d3030d269cc41..baf47d8eff349 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/FunctorCategory.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/FunctorCategory.lean @@ -54,14 +54,13 @@ work to convert to this version: namely, the natural isomorphism `(evaluation C D).obj k ⋙ prod.functor.obj (F.obj k) ≅ prod.functor.obj F ⋙ (evaluation C D).obj k` -/ -def FunctorCategory.prodPreservesColimits [HasBinaryProducts D] [HasColimits D] +lemma FunctorCategory.prod_preservesColimits [HasBinaryProducts D] [HasColimits D] [∀ X : D, PreservesColimits (prod.functor.obj X)] (F : C ⥤ D) : PreservesColimits (prod.functor.obj F) where preservesColimitsOfShape {J : Type u} [Category.{u, u} J] := { - preservesColimit := fun {K : J ⥤ C ⥤ D} => - ( { - preserves := fun {c : Cocone K} (t : IsColimit c) => by + preservesColimit := fun {K : J ⥤ C ⥤ D} => ({ + preserves := fun {c : Cocone K} (t : IsColimit c) => ⟨by apply evaluationJointlyReflectsColimits _ fun {k} => ?_ change IsColimit ((prod.functor.obj F ⋙ (evaluation _ _).obj k).mapCocone c) let this := @@ -71,7 +70,7 @@ def FunctorCategory.prodPreservesColimits [HasBinaryProducts D] [HasColimits D] · intro G apply asIso (prodComparison ((evaluation C D).obj k) F G) · intro G G' - apply prodComparison_natural ((evaluation C D).obj k) (𝟙 F) } ) } + apply prodComparison_natural ((evaluation C D).obj k) (𝟙 F)⟩ } ) } end @@ -81,43 +80,43 @@ variable {C : Type u₁} [Category.{v₁} C] variable {D : Type u₂} [Category.{v₂} D] variable {E : Type u₃} [Category.{v₃} E] -instance whiskeringLeftPreservesLimitsOfShape (J : Type u) [Category.{v} J] +instance whiskeringLeft_preservesLimitsOfShape (J : Type u) [Category.{v} J] [HasLimitsOfShape J D] (F : C ⥤ E) : PreservesLimitsOfShape J ((whiskeringLeft C E D).obj F) := ⟨fun {K} => - ⟨fun c {hc} => by + ⟨fun c {hc} => ⟨by apply evaluationJointlyReflectsLimits intro Y change IsLimit (((evaluation E D).obj (F.obj Y)).mapCone c) - exact PreservesLimit.preserves hc⟩⟩ + exact isLimitOfPreserves _ hc⟩⟩⟩ -instance whiskeringLeftPreservesColimitsOfShape (J : Type u) [Category.{v} J] +instance whiskeringLeft_preservesColimitsOfShape (J : Type u) [Category.{v} J] [HasColimitsOfShape J D] (F : C ⥤ E) : PreservesColimitsOfShape J ((whiskeringLeft C E D).obj F) := ⟨fun {K} => - ⟨fun c {hc} => by + ⟨fun c {hc} => ⟨by apply evaluationJointlyReflectsColimits intro Y change IsColimit (((evaluation E D).obj (F.obj Y)).mapCocone c) - exact PreservesColimit.preserves hc⟩⟩ + exact isColimitOfPreserves _ hc⟩⟩⟩ -instance whiskeringLeftPreservesLimits [HasLimitsOfSize.{w} D] (F : C ⥤ E) : +instance whiskeringLeft_preservesLimits [HasLimitsOfSize.{w, w'} D] (F : C ⥤ E) : PreservesLimitsOfSize.{w, w'} ((whiskeringLeft C E D).obj F) := - ⟨fun {J} _ => whiskeringLeftPreservesLimitsOfShape J F⟩ + ⟨fun {J} _ => whiskeringLeft_preservesLimitsOfShape J F⟩ -instance whiskeringLeftPreservesColimit [HasColimitsOfSize.{w} D] (F : C ⥤ E) : +instance whiskeringLeft_preservesColimit [HasColimitsOfSize.{w, w'} D] (F : C ⥤ E) : PreservesColimitsOfSize.{w, w'} ((whiskeringLeft C E D).obj F) := - ⟨fun {J} _ => whiskeringLeftPreservesColimitsOfShape J F⟩ + ⟨fun {J} _ => whiskeringLeft_preservesColimitsOfShape J F⟩ -instance whiskeringRightPreservesLimitsOfShape {C : Type*} [Category C] {D : Type*} +instance whiskeringRight_preservesLimitsOfShape {C : Type*} [Category C] {D : Type*} [Category D] {E : Type*} [Category E] {J : Type*} [Category J] [HasLimitsOfShape J D] (F : D ⥤ E) [PreservesLimitsOfShape J F] : PreservesLimitsOfShape J ((whiskeringRight C D E).obj F) := ⟨fun {K} => - ⟨fun c {hc} => by + ⟨fun c {hc} => ⟨by apply evaluationJointlyReflectsLimits _ (fun k => ?_) change IsLimit (((evaluation _ _).obj k ⋙ F).mapCone c) - exact PreservesLimit.preserves hc⟩⟩ + exact isLimitOfPreserves _ hc⟩⟩⟩ /-- Whiskering right and then taking a limit is the same as taking the limit and applying the functor. -/ @@ -143,15 +142,15 @@ theorem limitCompWhiskeringRightIsoLimitComp_hom_whiskerRight_π {C : Type*} [Ca limit.π (G ⋙ (whiskeringRight _ _ _).obj F) j := by simp [← Iso.eq_inv_comp] -instance whiskeringRightPreservesColimitsOfShape {C : Type*} [Category C] {D : Type*} +instance whiskeringRight_preservesColimitsOfShape {C : Type*} [Category C] {D : Type*} [Category D] {E : Type*} [Category E] {J : Type*} [Category J] [HasColimitsOfShape J D] (F : D ⥤ E) [PreservesColimitsOfShape J F] : PreservesColimitsOfShape J ((whiskeringRight C D E).obj F) := ⟨fun {K} => - ⟨fun c {hc} => by + ⟨fun c {hc} => ⟨by apply evaluationJointlyReflectsColimits _ (fun k => ?_) change IsColimit (((evaluation _ _).obj k ⋙ F).mapCocone c) - exact PreservesColimit.preserves hc⟩⟩ + exact isColimitOfPreserves _ hc⟩⟩⟩ /-- Whiskering right and then taking a colimit is the same as taking the colimit and applying the functor. -/ @@ -191,23 +190,23 @@ instance whiskeringRightPreservesColimits {C : Type*} [Category C] {D : Type*} [ -- Porting note: fixed spelling mistake in def /-- If `Lan F.op : (Cᵒᵖ ⥤ Type*) ⥤ (Dᵒᵖ ⥤ Type*)` preserves limits of shape `J`, so will `F`. -/ -noncomputable def preservesLimitOfLanPreservesLimit {C D : Type u} [SmallCategory C] +lemma preservesLimit_of_lan_preservesLimit {C D : Type u} [SmallCategory C] [SmallCategory D] (F : C ⥤ D) (J : Type u) [SmallCategory J] [PreservesLimitsOfShape J (F.op.lan : _ ⥤ Dᵒᵖ ⥤ Type u)] : PreservesLimitsOfShape J F := by - apply @preservesLimitsOfShapeOfReflectsOfPreserves _ _ _ _ _ _ _ _ F yoneda ?_ - exact preservesLimitsOfShapeOfNatIso (Presheaf.compYonedaIsoYonedaCompLan F).symm + apply @preservesLimitsOfShape_of_reflects_of_preserves _ _ _ _ _ _ _ _ F yoneda ?_ + exact preservesLimitsOfShape_of_natIso (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] +lemma preservesFiniteLimits_of_evaluation {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 _⟩ + ⟨fun J _ _ => preservesLimitsOfShape_of_evaluation 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] +lemma preservesFiniteColimits_of_evaluation {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 _⟩ + ⟨fun J _ _ => preservesColimitsOfShape_of_evaluation F J fun k => (h k).preservesFiniteColimits _⟩ end @@ -224,14 +223,14 @@ variable [HasLimitsOfShape J C] [HasColimitsOfShape K C] variable [PreservesLimitsOfShape J (colim : (K ⥤ C) ⥤ _)] noncomputable instance : PreservesLimitsOfShape J (colim : (K ⥤ D ⥤ C) ⥤ _) := - preservesLimitsOfShapeOfEvaluation _ _ (fun d => + preservesLimitsOfShape_of_evaluation _ _ (fun d => let i : (colim : (K ⥤ D ⥤ C) ⥤ _) ⋙ (evaluation D C).obj d ≅ colimit ((whiskeringRight K (D ⥤ C) C).obj ((evaluation D C).obj d)).flip := NatIso.ofComponents (fun X => (colimitObjIsoColimitCompEvaluation _ _) ≪≫ (by exact HasColimit.isoOfNatIso (Iso.refl _)) ≪≫ (colimitObjIsoColimitCompEvaluation _ _).symm) (fun {F G} η => colimit_obj_ext (fun j => by simp [← NatTrans.comp_app_assoc])) - preservesLimitsOfShapeOfNatIso (i ≪≫ colimitFlipIsoCompColim _).symm) + preservesLimitsOfShape_of_natIso (i ≪≫ colimitFlipIsoCompColim _).symm) end @@ -241,14 +240,14 @@ variable [HasColimitsOfShape J C] [HasLimitsOfShape K C] variable [PreservesColimitsOfShape J (lim : (K ⥤ C) ⥤ _)] noncomputable instance : PreservesColimitsOfShape J (lim : (K ⥤ D ⥤ C) ⥤ _) := - preservesColimitsOfShapeOfEvaluation _ _ (fun d => + preservesColimitsOfShape_of_evaluation _ _ (fun d => let i : (lim : (K ⥤ D ⥤ C) ⥤ _) ⋙ (evaluation D C).obj d ≅ limit ((whiskeringRight K (D ⥤ C) C).obj ((evaluation D C).obj d)).flip := NatIso.ofComponents (fun X => (limitObjIsoLimitCompEvaluation _ _) ≪≫ (by exact HasLimit.isoOfNatIso (Iso.refl _)) ≪≫ (limitObjIsoLimitCompEvaluation _ _).symm) (fun {F G} η => limit_obj_ext (fun j => by simp [← NatTrans.comp_app])) - preservesColimitsOfShapeOfNatIso (i ≪≫ limitFlipIsoCompLim _).symm) + preservesColimitsOfShape_of_natIso (i ≪≫ limitFlipIsoCompLim _).symm) end diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Limits.lean b/Mathlib/CategoryTheory/Limits/Preserves/Limits.lean index ac77b2c0bcd51..f0517bf3e7d4b 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Limits.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Limits.lean @@ -38,8 +38,8 @@ variable [PreservesLimit F G] @[simp] theorem preserves_lift_mapCone (c₁ c₂ : Cone F) (t : IsLimit c₁) : - (PreservesLimit.preserves t).lift (G.mapCone c₂) = G.map (t.lift c₂) := - ((PreservesLimit.preserves t).uniq (G.mapCone c₂) _ (by simp [← G.map_comp])).symm + (isLimitOfPreserves G t).lift (G.mapCone c₂) = G.map (t.lift c₂) := + ((isLimitOfPreserves G t).uniq (G.mapCone c₂) _ (by simp [← G.map_comp])).symm variable [HasLimit F] @@ -47,7 +47,7 @@ variable [HasLimit F] to the limit of the functor `F ⋙ G`. -/ def preservesLimitIso : G.obj (limit F) ≅ limit (F ⋙ G) := - (PreservesLimit.preserves (limit.isLimit _)).conePointUniqueUpToIso (limit.isLimit _) + (isLimitOfPreserves G (limit.isLimit _)).conePointUniqueUpToIso (limit.isLimit _) @[reassoc (attr := simp)] theorem preservesLimitIso_hom_π (j) : @@ -98,8 +98,8 @@ variable [HasLimit F] [HasLimit (F ⋙ G)] /-- If the comparison morphism `G.obj (limit F) ⟶ limit (F ⋙ G)` is an isomorphism, then `G` preserves limits of `F`. -/ -def preservesLimitOfIsIsoPost [IsIso (limit.post F G)] : PreservesLimit F G := - preservesLimitOfPreservesLimitCone (limit.isLimit F) (by +lemma preservesLimit_of_isIso_post [IsIso (limit.post F G)] : PreservesLimit F G := + preservesLimit_of_preserves_limit_cone (limit.isLimit F) (by convert IsLimit.ofPointIso (limit.isLimit (F ⋙ G)) assumption) @@ -111,8 +111,8 @@ variable [PreservesColimit F G] @[simp] theorem preserves_desc_mapCocone (c₁ c₂ : Cocone F) (t : IsColimit c₁) : - (PreservesColimit.preserves t).desc (G.mapCocone _) = G.map (t.desc c₂) := - ((PreservesColimit.preserves t).uniq (G.mapCocone _) _ (by simp [← G.map_comp])).symm + (isColimitOfPreserves G t).desc (G.mapCocone _) = G.map (t.desc c₂) := + ((isColimitOfPreserves G t).uniq (G.mapCocone _) _ (by simp [← G.map_comp])).symm variable [HasColimit F] @@ -121,7 +121,7 @@ variable [HasColimit F] to the colimit of the functor `F ⋙ G`. -/ def preservesColimitIso : G.obj (colimit F) ≅ colimit (F ⋙ G) := - (PreservesColimit.preserves (colimit.isColimit _)).coconePointUniqueUpToIso (colimit.isColimit _) + (isColimitOfPreserves G (colimit.isColimit _)).coconePointUniqueUpToIso (colimit.isColimit _) @[reassoc (attr := simp)] theorem ι_preservesColimitIso_inv (j : J) : @@ -134,7 +134,7 @@ alias ι_preservesColimitsIso_inv := ι_preservesColimitIso_inv @[reassoc (attr := simp)] theorem ι_preservesColimitIso_hom (j : J) : G.map (colimit.ι F j) ≫ (preservesColimitIso G F).hom = colimit.ι (F ⋙ G) j := - (PreservesColimit.preserves (colimit.isColimit _)).comp_coconePointUniqueUpToIso_hom _ j + (isColimitOfPreserves G (colimit.isColimit _)).comp_coconePointUniqueUpToIso_hom _ j @[deprecated (since := "2024-10-27")] alias ι_preservesColimitsIso_hom := ι_preservesColimitIso_hom @@ -177,8 +177,8 @@ variable [HasColimit F] [HasColimit (F ⋙ G)] /-- If the comparison morphism `colimit (F ⋙ G) ⟶ G.obj (colimit F)` is an isomorphism, then `G` preserves colimits of `F`. -/ -def preservesColimitOfIsIsoPost [IsIso (colimit.post F G)] : PreservesColimit F G := - preservesColimitOfPreservesColimitCocone (colimit.isColimit F) (by +lemma preservesColimit_of_isIso_post [IsIso (colimit.post F G)] : PreservesColimit F G := + preservesColimit_of_preserves_colimit_cocone (colimit.isColimit F) (by convert IsColimit.ofPointIso (colimit.isColimit (F ⋙ G)) assumption) diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Opposites.lean b/Mathlib/CategoryTheory/Limits/Preserves/Opposites.lean index 99f59eb1dea7a..3c6706cf053d9 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Opposites.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Opposites.lean @@ -28,112 +28,113 @@ variable {J : Type w} [Category.{w'} J] /-- If `F : C ⥤ D` preserves colimits of `K.leftOp : Jᵒᵖ ⥤ C`, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits of `K : J ⥤ Cᵒᵖ`. -/ -def preservesLimitOp (K : J ⥤ Cᵒᵖ) (F : C ⥤ D) [PreservesColimit K.leftOp F] : +lemma preservesLimit_op (K : J ⥤ Cᵒᵖ) (F : C ⥤ D) [PreservesColimit K.leftOp F] : PreservesLimit K F.op where preserves {_} hc := - isLimitConeRightOpOfCocone _ (isColimitOfPreserves F (isColimitCoconeLeftOpOfCone _ hc)) + ⟨isLimitConeRightOpOfCocone _ (isColimitOfPreserves F (isColimitCoconeLeftOpOfCone _ hc))⟩ /-- If `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits of `K.op : Jᵒᵖ ⥤ Cᵒᵖ`, then `F : C ⥤ D` preserves limits of `K : J ⥤ C`. -/ -def preservesLimitOfOp (K : J ⥤ C) (F : C ⥤ D) [PreservesColimit K.op F.op] : +lemma preservesLimit_of_op (K : J ⥤ C) (F : C ⥤ D) [PreservesColimit K.op F.op] : PreservesLimit K F where - preserves {_} hc := isLimitOfOp (isColimitOfPreserves F.op (IsLimit.op hc)) + preserves {_} hc := ⟨isLimitOfOp (isColimitOfPreserves F.op (IsLimit.op hc))⟩ /-- If `F : C ⥤ Dᵒᵖ` preserves colimits of `K.leftOp : Jᵒᵖ ⥤ C`, then `F.leftOp : Cᵒᵖ ⥤ D` preserves limits of `K : J ⥤ Cᵒᵖ`. -/ -def preservesLimitLeftOp (K : J ⥤ Cᵒᵖ) (F : C ⥤ Dᵒᵖ) [PreservesColimit K.leftOp F] : +lemma preservesLimit_leftOp (K : J ⥤ Cᵒᵖ) (F : C ⥤ Dᵒᵖ) [PreservesColimit K.leftOp F] : PreservesLimit K F.leftOp where preserves {_} hc := - isLimitConeUnopOfCocone _ (isColimitOfPreserves F (isColimitCoconeLeftOpOfCone _ hc)) + ⟨isLimitConeUnopOfCocone _ (isColimitOfPreserves F (isColimitCoconeLeftOpOfCone _ hc))⟩ /-- If `F.leftOp : Cᵒᵖ ⥤ D` preserves colimits of `K.op : Jᵒᵖ ⥤ Cᵒᵖ`, then `F : C ⥤ Dᵒᵖ` preserves limits of `K : J ⥤ C`. -/ -def preservesLimitOfLeftOp (K : J ⥤ C) (F : C ⥤ Dᵒᵖ) [PreservesColimit K.op F.leftOp] : +lemma preservesLimit_of_leftOp (K : J ⥤ C) (F : C ⥤ Dᵒᵖ) [PreservesColimit K.op F.leftOp] : PreservesLimit K F where - preserves {_} hc := isLimitOfCoconeLeftOpOfCone _ (isColimitOfPreserves F.leftOp (IsLimit.op hc)) + preserves {_} hc := + ⟨isLimitOfCoconeLeftOpOfCone _ (isColimitOfPreserves F.leftOp (IsLimit.op hc))⟩ /-- If `F : Cᵒᵖ ⥤ D` preserves colimits of `K.op : Jᵒᵖ ⥤ Cᵒᵖ`, then `F.rightOp : C ⥤ Dᵒᵖ` preserves limits of `K : J ⥤ C`. -/ -def preservesLimitRightOp (K : J ⥤ C) (F : Cᵒᵖ ⥤ D) [PreservesColimit K.op F] : +lemma preservesLimit_rightOp (K : J ⥤ C) (F : Cᵒᵖ ⥤ D) [PreservesColimit K.op F] : PreservesLimit K F.rightOp where preserves {_} hc := - isLimitConeRightOpOfCocone _ (isColimitOfPreserves F hc.op) + ⟨isLimitConeRightOpOfCocone _ (isColimitOfPreserves F hc.op)⟩ /-- If `F.rightOp : C ⥤ Dᵒᵖ` preserves colimits of `K.leftOp : Jᵒᵖ ⥤ Cᵒᵖ`, then `F : Cᵒᵖ ⥤ D` preserves limits of `K : J ⥤ Cᵒᵖ`. -/ -def preservesLimitOfRightOp (K : J ⥤ Cᵒᵖ) (F : Cᵒᵖ ⥤ D) [PreservesColimit K.leftOp F.rightOp] : +lemma preservesLimit_of_rightOp (K : J ⥤ Cᵒᵖ) (F : Cᵒᵖ ⥤ D) [PreservesColimit K.leftOp F.rightOp] : PreservesLimit K F where preserves {_} hc := - isLimitOfOp (isColimitOfPreserves F.rightOp (isColimitCoconeLeftOpOfCone _ hc)) + ⟨isLimitOfOp (isColimitOfPreserves F.rightOp (isColimitCoconeLeftOpOfCone _ hc))⟩ /-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits of `K.op : Jᵒᵖ ⥤ Cᵒᵖ`, then `F.unop : C ⥤ D` preserves limits of `K : J ⥤ C`. -/ -def preservesLimitUnop (K : J ⥤ C) (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimit K.op F] : +lemma preservesLimit_unop (K : J ⥤ C) (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimit K.op F] : PreservesLimit K F.unop where preserves {_} hc := - isLimitConeUnopOfCocone _ (isColimitOfPreserves F hc.op) + ⟨isLimitConeUnopOfCocone _ (isColimitOfPreserves F hc.op)⟩ /-- If `F.unop : C ⥤ D` preserves colimits of `K.leftOp : Jᵒᵖ ⥤ C`, then `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits of `K : J ⥤ Cᵒᵖ`. -/ -def preservesLimitOfUnop (K : J ⥤ Cᵒᵖ) (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimit K.leftOp F.unop] : +lemma preservesLimit_of_unop (K : J ⥤ Cᵒᵖ) (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimit K.leftOp F.unop] : PreservesLimit K F where preserves {_} hc := - isLimitOfCoconeLeftOpOfCone _ (isColimitOfPreserves F.unop (isColimitCoconeLeftOpOfCone _ hc)) + ⟨isLimitOfCoconeLeftOpOfCone _ (isColimitOfPreserves F.unop (isColimitCoconeLeftOpOfCone _ hc))⟩ /-- If `F : C ⥤ D` preserves limits of `K.leftOp : Jᵒᵖ ⥤ C`, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits of `K : J ⥤ Cᵒᵖ`. -/ -def preservesColimitOp (K : J ⥤ Cᵒᵖ) (F : C ⥤ D) [PreservesLimit K.leftOp F] : +lemma preservesColimit_op (K : J ⥤ Cᵒᵖ) (F : C ⥤ D) [PreservesLimit K.leftOp F] : PreservesColimit K F.op where preserves {_} hc := - isColimitCoconeRightOpOfCone _ (isLimitOfPreserves F (isLimitConeLeftOpOfCocone _ hc)) + ⟨isColimitCoconeRightOpOfCone _ (isLimitOfPreserves F (isLimitConeLeftOpOfCocone _ hc))⟩ /-- If `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits of `K.op : Jᵒᵖ ⥤ Cᵒᵖ`, then `F : C ⥤ D` preserves colimits of `K : J ⥤ C`. -/ -def preservesColimitOfOp (K : J ⥤ C) (F : C ⥤ D) [PreservesLimit K.op F.op] : +lemma preservesColimit_of_op (K : J ⥤ C) (F : C ⥤ D) [PreservesLimit K.op F.op] : PreservesColimit K F where - preserves {_} hc := isColimitOfOp (isLimitOfPreserves F.op (IsColimit.op hc)) + preserves {_} hc := ⟨isColimitOfOp (isLimitOfPreserves F.op (IsColimit.op hc))⟩ /-- If `F : C ⥤ Dᵒᵖ` preserves limits of `K.leftOp : Jᵒᵖ ⥤ C`, then `F.leftOp : Cᵒᵖ ⥤ D` preserves colimits of `K : J ⥤ Cᵒᵖ`. -/ -def preservesColimitLeftOp (K : J ⥤ Cᵒᵖ) (F : C ⥤ Dᵒᵖ) [PreservesLimit K.leftOp F] : +lemma preservesColimit_leftOp (K : J ⥤ Cᵒᵖ) (F : C ⥤ Dᵒᵖ) [PreservesLimit K.leftOp F] : PreservesColimit K F.leftOp where preserves {_} hc := - isColimitCoconeUnopOfCone _ (isLimitOfPreserves F (isLimitConeLeftOpOfCocone _ hc)) + ⟨isColimitCoconeUnopOfCone _ (isLimitOfPreserves F (isLimitConeLeftOpOfCocone _ hc))⟩ /-- If `F.leftOp : Cᵒᵖ ⥤ D` preserves limits of `K.op : Jᵒᵖ ⥤ Cᵒᵖ`, then `F : C ⥤ Dᵒᵖ` preserves colimits of `K : J ⥤ C`. -/ -def preservesColimitOfLeftOp (K : J ⥤ C) (F : C ⥤ Dᵒᵖ) [PreservesLimit K.op F.leftOp] : +lemma preservesColimit_of_leftOp (K : J ⥤ C) (F : C ⥤ Dᵒᵖ) [PreservesLimit K.op F.leftOp] : PreservesColimit K F where preserves {_} hc := - isColimitOfConeLeftOpOfCocone _ (isLimitOfPreserves F.leftOp (IsColimit.op hc)) + ⟨isColimitOfConeLeftOpOfCocone _ (isLimitOfPreserves F.leftOp (IsColimit.op hc))⟩ /-- If `F : Cᵒᵖ ⥤ D` preserves limits of `K.op : Jᵒᵖ ⥤ Cᵒᵖ`, then `F.rightOp : C ⥤ Dᵒᵖ` preserves colimits of `K : J ⥤ C`. -/ -def preservesColimitRightOp (K : J ⥤ C) (F : Cᵒᵖ ⥤ D) [PreservesLimit K.op F] : +lemma preservesColimit_rightOp (K : J ⥤ C) (F : Cᵒᵖ ⥤ D) [PreservesLimit K.op F] : PreservesColimit K F.rightOp where preserves {_} hc := - isColimitCoconeRightOpOfCone _ (isLimitOfPreserves F hc.op) + ⟨isColimitCoconeRightOpOfCone _ (isLimitOfPreserves F hc.op)⟩ /-- If `F.rightOp : C ⥤ Dᵒᵖ` preserves limits of `K.leftOp : Jᵒᵖ ⥤ C`, then `F : Cᵒᵖ ⥤ D` preserves colimits of `K : J ⥤ Cᵒᵖ`. -/ -def preservesColimitOfRightOp (K : J ⥤ Cᵒᵖ) (F : Cᵒᵖ ⥤ D) [PreservesLimit K.leftOp F.rightOp] : +lemma preservesColimit_of_rightOp (K : J ⥤ Cᵒᵖ) (F : Cᵒᵖ ⥤ D) [PreservesLimit K.leftOp F.rightOp] : PreservesColimit K F where preserves {_} hc := - isColimitOfOp (isLimitOfPreserves F.rightOp (isLimitConeLeftOpOfCocone _ hc)) + ⟨isColimitOfOp (isLimitOfPreserves F.rightOp (isLimitConeLeftOpOfCocone _ hc))⟩ /-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits of `K.op : Jᵒᵖ ⥤ Cᵒᵖ`, then `F.unop : C ⥤ D` preserves colimits of `K : J ⥤ C`. -/ -def preservesColimitUnop (K : J ⥤ C) (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimit K.op F] : +lemma preservesColimit_unop (K : J ⥤ C) (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimit K.op F] : PreservesColimit K F.unop where preserves {_} hc := - isColimitCoconeUnopOfCone _ (isLimitOfPreserves F hc.op) + ⟨isColimitCoconeUnopOfCone _ (isLimitOfPreserves F hc.op)⟩ /-- If `F.unop : C ⥤ D` preserves limits of `K.op : Jᵒᵖ ⥤ C`, then `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits of `K : J ⥤ Cᵒᵖ`. -/ -def preservesColimitOfUnop (K : J ⥤ Cᵒᵖ) (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimit K.leftOp F.unop] : +lemma preservesColimit_of_unop (K : J ⥤ Cᵒᵖ) (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimit K.leftOp F.unop] : PreservesColimit K F where preserves {_} hc := - isColimitOfConeLeftOpOfCocone _ (isLimitOfPreserves F.unop (isLimitConeLeftOpOfCocone _ hc)) + ⟨isColimitOfConeLeftOpOfCocone _ (isLimitOfPreserves F.unop (isLimitConeLeftOpOfCocone _ hc))⟩ section @@ -141,385 +142,388 @@ variable (J) /-- If `F : C ⥤ D` preserves colimits of shape `Jᵒᵖ`, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits of shape `J`. -/ -def preservesLimitsOfShapeOp (F : C ⥤ D) [PreservesColimitsOfShape Jᵒᵖ F] : - PreservesLimitsOfShape J F.op where preservesLimit {K} := preservesLimitOp K F +lemma preservesLimitsOfShape_op (F : C ⥤ D) [PreservesColimitsOfShape Jᵒᵖ F] : + PreservesLimitsOfShape J F.op where preservesLimit {K} := preservesLimit_op K F /-- If `F : C ⥤ Dᵒᵖ` preserves colimits of shape `Jᵒᵖ`, then `F.leftOp : Cᵒᵖ ⥤ D` preserves limits of shape `J`. -/ -def preservesLimitsOfShapeLeftOp (F : C ⥤ Dᵒᵖ) [PreservesColimitsOfShape Jᵒᵖ F] : - PreservesLimitsOfShape J F.leftOp where preservesLimit {K} := preservesLimitLeftOp K F +lemma preservesLimitsOfShape_leftOp (F : C ⥤ Dᵒᵖ) [PreservesColimitsOfShape Jᵒᵖ F] : + PreservesLimitsOfShape J F.leftOp where preservesLimit {K} := preservesLimit_leftOp K F /-- If `F : Cᵒᵖ ⥤ D` preserves colimits of shape `Jᵒᵖ`, then `F.rightOp : C ⥤ Dᵒᵖ` preserves limits of shape `J`. -/ -def preservesLimitsOfShapeRightOp (F : Cᵒᵖ ⥤ D) [PreservesColimitsOfShape Jᵒᵖ F] : - PreservesLimitsOfShape J F.rightOp where preservesLimit {K} := preservesLimitRightOp K F +lemma preservesLimitsOfShape_rightOp (F : Cᵒᵖ ⥤ D) [PreservesColimitsOfShape Jᵒᵖ F] : + PreservesLimitsOfShape J F.rightOp where preservesLimit {K} := preservesLimit_rightOp K F /-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits of shape `Jᵒᵖ`, then `F.unop : C ⥤ D` preserves limits of shape `J`. -/ -def preservesLimitsOfShapeUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimitsOfShape Jᵒᵖ F] : - PreservesLimitsOfShape J F.unop where preservesLimit {K} := preservesLimitUnop K F +lemma preservesLimitsOfShape_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimitsOfShape Jᵒᵖ F] : + PreservesLimitsOfShape J F.unop where preservesLimit {K} := preservesLimit_unop K F /-- If `F : C ⥤ D` preserves limits of shape `Jᵒᵖ`, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits of shape `J`. -/ -def preservesColimitsOfShapeOp (F : C ⥤ D) [PreservesLimitsOfShape Jᵒᵖ F] : - PreservesColimitsOfShape J F.op where preservesColimit {K} := preservesColimitOp K F +lemma preservesColimitsOfShape_op (F : C ⥤ D) [PreservesLimitsOfShape Jᵒᵖ F] : + PreservesColimitsOfShape J F.op where preservesColimit {K} := preservesColimit_op K F /-- If `F : C ⥤ Dᵒᵖ` preserves limits of shape `Jᵒᵖ`, then `F.leftOp : Cᵒᵖ ⥤ D` preserves colimits of shape `J`. -/ -def preservesColimitsOfShapeLeftOp (F : C ⥤ Dᵒᵖ) [PreservesLimitsOfShape Jᵒᵖ F] : - PreservesColimitsOfShape J F.leftOp where preservesColimit {K} := preservesColimitLeftOp K F +lemma preservesColimitsOfShape_leftOp (F : C ⥤ Dᵒᵖ) [PreservesLimitsOfShape Jᵒᵖ F] : + PreservesColimitsOfShape J F.leftOp where preservesColimit {K} := preservesColimit_leftOp K F /-- If `F : Cᵒᵖ ⥤ D` preserves limits of shape `Jᵒᵖ`, then `F.rightOp : C ⥤ Dᵒᵖ` preserves colimits of shape `J`. -/ -def preservesColimitsOfShapeRightOp (F : Cᵒᵖ ⥤ D) [PreservesLimitsOfShape Jᵒᵖ F] : - PreservesColimitsOfShape J F.rightOp where preservesColimit {K} := preservesColimitRightOp K F +lemma preservesColimitsOfShape_rightOp (F : Cᵒᵖ ⥤ D) [PreservesLimitsOfShape Jᵒᵖ F] : + PreservesColimitsOfShape J F.rightOp where preservesColimit {K} := preservesColimit_rightOp K F /-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits of shape `Jᵒᵖ`, then `F.unop : C ⥤ D` preserves colimits of shape `J`. -/ -def preservesColimitsOfShapeUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimitsOfShape Jᵒᵖ F] : - PreservesColimitsOfShape J F.unop where preservesColimit {K} := preservesColimitUnop K F +lemma preservesColimitsOfShape_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimitsOfShape Jᵒᵖ F] : + PreservesColimitsOfShape J F.unop where preservesColimit {K} := preservesColimit_unop K F /-- If `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits of shape `Jᵒᵖ`, then `F : C ⥤ D` preserves limits of shape `J`. -/ -def preservesLimitsOfShapeOfOp (F : C ⥤ D) [PreservesColimitsOfShape Jᵒᵖ F.op] : - PreservesLimitsOfShape J F where preservesLimit {K} := preservesLimitOfOp K F +lemma preservesLimitsOfShape_of_op (F : C ⥤ D) [PreservesColimitsOfShape Jᵒᵖ F.op] : + PreservesLimitsOfShape J F where preservesLimit {K} := preservesLimit_of_op K F /-- If `F.leftOp : Cᵒᵖ ⥤ D` preserves colimits of shape `Jᵒᵖ`, then `F : C ⥤ Dᵒᵖ` preserves limits of shape `J`. -/ -def preservesLimitsOfShapeOfLeftOp (F : C ⥤ Dᵒᵖ) [PreservesColimitsOfShape Jᵒᵖ F.leftOp] : - PreservesLimitsOfShape J F where preservesLimit {K} := preservesLimitOfLeftOp K F +lemma preservesLimitsOfShape_of_leftOp (F : C ⥤ Dᵒᵖ) [PreservesColimitsOfShape Jᵒᵖ F.leftOp] : + PreservesLimitsOfShape J F where preservesLimit {K} := preservesLimit_of_leftOp K F /-- If `F.rightOp : C ⥤ Dᵒᵖ` preserves colimits of shape `Jᵒᵖ`, then `F : Cᵒᵖ ⥤ D` preserves limits of shape `J`. -/ -def preservesLimitsOfShapeOfRightOp (F : Cᵒᵖ ⥤ D) [PreservesColimitsOfShape Jᵒᵖ F.rightOp] : - PreservesLimitsOfShape J F where preservesLimit {K} := preservesLimitOfRightOp K F +lemma preservesLimitsOfShape_of_rightOp (F : Cᵒᵖ ⥤ D) [PreservesColimitsOfShape Jᵒᵖ F.rightOp] : + PreservesLimitsOfShape J F where preservesLimit {K} := preservesLimit_of_rightOp K F /-- If `F.unop : C ⥤ D` preserves colimits of shape `Jᵒᵖ`, then `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits of shape `J`. -/ -def preservesLimitsOfShapeOfUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimitsOfShape Jᵒᵖ F.unop] : - PreservesLimitsOfShape J F where preservesLimit {K} := preservesLimitOfUnop K F +lemma preservesLimitsOfShape_of_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimitsOfShape Jᵒᵖ F.unop] : + PreservesLimitsOfShape J F where preservesLimit {K} := preservesLimit_of_unop K F /-- If `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits of shape `Jᵒᵖ`, then `F : C ⥤ D` preserves colimits of shape `J`. -/ -def preservesColimitsOfShapeOfOp (F : C ⥤ D) [PreservesLimitsOfShape Jᵒᵖ F.op] : - PreservesColimitsOfShape J F where preservesColimit {K} := preservesColimitOfOp K F +lemma preservesColimitsOfShape_of_op (F : C ⥤ D) [PreservesLimitsOfShape Jᵒᵖ F.op] : + PreservesColimitsOfShape J F where preservesColimit {K} := preservesColimit_of_op K F /-- If `F.leftOp : Cᵒᵖ ⥤ D` preserves limits of shape `Jᵒᵖ`, then `F : C ⥤ Dᵒᵖ` preserves colimits of shape `J`. -/ -def preservesColimitsOfShapeOfLeftOp (F : C ⥤ Dᵒᵖ) [PreservesLimitsOfShape Jᵒᵖ F.leftOp] : - PreservesColimitsOfShape J F where preservesColimit {K} := preservesColimitOfLeftOp K F +lemma preservesColimitsOfShape_of_leftOp (F : C ⥤ Dᵒᵖ) [PreservesLimitsOfShape Jᵒᵖ F.leftOp] : + PreservesColimitsOfShape J F where preservesColimit {K} := preservesColimit_of_leftOp K F /-- If `F.rightOp : C ⥤ Dᵒᵖ` preserves limits of shape `Jᵒᵖ`, then `F : Cᵒᵖ ⥤ D` preserves colimits of shape `J`. -/ -def preservesColimitsOfShapeOfRightOp (F : Cᵒᵖ ⥤ D) [PreservesLimitsOfShape Jᵒᵖ F.rightOp] : - PreservesColimitsOfShape J F where preservesColimit {K} := preservesColimitOfRightOp K F +lemma preservesColimitsOfShape_of_rightOp (F : Cᵒᵖ ⥤ D) [PreservesLimitsOfShape Jᵒᵖ F.rightOp] : + PreservesColimitsOfShape J F where preservesColimit {K} := preservesColimit_of_rightOp K F /-- If `F.unop : C ⥤ D` preserves limits of shape `Jᵒᵖ`, then `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits of shape `J`. -/ -def preservesColimitsOfShapeOfUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimitsOfShape Jᵒᵖ F.unop] : - PreservesColimitsOfShape J F where preservesColimit {K} := preservesColimitOfUnop K F +lemma preservesColimitsOfShape_of_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimitsOfShape Jᵒᵖ F.unop] : + PreservesColimitsOfShape J F where preservesColimit {K} := preservesColimit_of_unop K F end /-- If `F : C ⥤ D` preserves colimits, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits. -/ -def preservesLimitsOfSizeOp (F : C ⥤ D) [PreservesColimitsOfSize.{w, w'} F] : +lemma preservesLimitsOfSize_op (F : C ⥤ D) [PreservesColimitsOfSize.{w, w'} F] : PreservesLimitsOfSize.{w, w'} F.op where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeOp _ _ + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_op _ _ /-- If `F : C ⥤ Dᵒᵖ` preserves colimits, then `F.leftOp : Cᵒᵖ ⥤ D` preserves limits. -/ -def preservesLimitsOfSizeLeftOp (F : C ⥤ Dᵒᵖ) [PreservesColimitsOfSize.{w, w'} F] : +lemma preservesLimitsOfSize_leftOp (F : C ⥤ Dᵒᵖ) [PreservesColimitsOfSize.{w, w'} F] : PreservesLimitsOfSize.{w, w'} F.leftOp where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeLeftOp _ _ + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_leftOp _ _ /-- If `F : Cᵒᵖ ⥤ D` preserves colimits, then `F.rightOp : C ⥤ Dᵒᵖ` preserves limits. -/ -def preservesLimitsOfSizeRightOp (F : Cᵒᵖ ⥤ D) [PreservesColimitsOfSize.{w, w'} F] : +lemma preservesLimitsOfSize_rightOp (F : Cᵒᵖ ⥤ D) [PreservesColimitsOfSize.{w, w'} F] : PreservesLimitsOfSize.{w, w'} F.rightOp where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeRightOp _ _ + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_rightOp _ _ /-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits, then `F.unop : C ⥤ D` preserves limits. -/ -def preservesLimitsOfSizeUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimitsOfSize.{w, w'} F] : +lemma preservesLimitsOfSize_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimitsOfSize.{w, w'} F] : PreservesLimitsOfSize.{w, w'} F.unop where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeUnop _ _ + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_unop _ _ /-- If `F : C ⥤ D` preserves limits, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits. -/ -def preservesColimitsOfSizeOp (F : C ⥤ D) [PreservesLimitsOfSize.{w, w'} F] : +lemma preservesColimitsOfSize_op (F : C ⥤ D) [PreservesLimitsOfSize.{w, w'} F] : PreservesColimitsOfSize.{w, w'} F.op where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeOp _ _ + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_op _ _ /-- If `F : C ⥤ Dᵒᵖ` preserves limits, then `F.leftOp : Cᵒᵖ ⥤ D` preserves colimits. -/ -def preservesColimitsOfSizeLeftOp (F : C ⥤ Dᵒᵖ) [PreservesLimitsOfSize.{w, w'} F] : +lemma preservesColimitsOfSize_leftOp (F : C ⥤ Dᵒᵖ) [PreservesLimitsOfSize.{w, w'} F] : PreservesColimitsOfSize.{w, w'} F.leftOp where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeLeftOp _ _ + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_leftOp _ _ /-- If `F : Cᵒᵖ ⥤ D` preserves limits, then `F.rightOp : C ⥤ Dᵒᵖ` preserves colimits. -/ -def preservesColimitsOfSizeRightOp (F : Cᵒᵖ ⥤ D) [PreservesLimitsOfSize.{w, w'} F] : +lemma preservesColimitsOfSize_rightOp (F : Cᵒᵖ ⥤ D) [PreservesLimitsOfSize.{w, w'} F] : PreservesColimitsOfSize.{w, w'} F.rightOp where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeRightOp _ _ + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_rightOp _ _ /-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits, then `F.unop : C ⥤ D` preserves colimits. -/ -def preservesColimitsOfSizeUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimitsOfSize.{w, w'} F] : +lemma preservesColimitsOfSize_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimitsOfSize.{w, w'} F] : PreservesColimitsOfSize.{w, w'} F.unop where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeUnop _ _ + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_unop _ _ /-- If `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits, then `F : C ⥤ D` preserves limits. -/ -def preservesLimitsOfSizeOfOp (F : C ⥤ D) [PreservesColimitsOfSize.{w, w'} F.op] : +lemma preservesLimitsOfSize_of_op (F : C ⥤ D) [PreservesColimitsOfSize.{w, w'} F.op] : PreservesLimitsOfSize.{w, w'} F where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeOfOp _ _ + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_of_op _ _ /-- If `F.leftOp : Cᵒᵖ ⥤ D` preserves colimits, then `F : C ⥤ Dᵒᵖ` preserves limits. -/ -def preservesLimitsOfSizeOfLeftOp (F : C ⥤ Dᵒᵖ) [PreservesColimitsOfSize.{w, w'} F.leftOp] : +lemma preservesLimitsOfSize_of_leftOp (F : C ⥤ Dᵒᵖ) [PreservesColimitsOfSize.{w, w'} F.leftOp] : PreservesLimitsOfSize.{w, w'} F where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeOfLeftOp _ _ + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_of_leftOp _ _ /-- If `F.rightOp : C ⥤ Dᵒᵖ` preserves colimits, then `F : Cᵒᵖ ⥤ D` preserves limits. -/ -def preservesLimitsOfSizeOfRightOp (F : Cᵒᵖ ⥤ D) [PreservesColimitsOfSize.{w, w'} F.rightOp] : +lemma preservesLimitsOfSize_of_rightOp (F : Cᵒᵖ ⥤ D) [PreservesColimitsOfSize.{w, w'} F.rightOp] : PreservesLimitsOfSize.{w, w'} F where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeOfRightOp _ _ + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_of_rightOp _ _ /-- If `F.unop : C ⥤ D` preserves colimits, then `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits. -/ -def preservesLimitsOfSizeOfUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimitsOfSize.{w, w'} F.unop] : +lemma preservesLimitsOfSize_of_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimitsOfSize.{w, w'} F.unop] : PreservesLimitsOfSize.{w, w'} F where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeOfUnop _ _ + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_of_unop _ _ /-- If `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits, then `F : C ⥤ D` preserves colimits. -/ -def preservesColimitsOfSizeOfOp (F : C ⥤ D) [PreservesLimitsOfSize.{w, w'} F.op] : +lemma preservesColimitsOfSize_of_op (F : C ⥤ D) [PreservesLimitsOfSize.{w, w'} F.op] : PreservesColimitsOfSize.{w, w'} F where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeOfOp _ _ + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_of_op _ _ /-- If `F.leftOp : Cᵒᵖ ⥤ D` preserves limits, then `F : C ⥤ Dᵒᵖ` preserves colimits. -/ -def preservesColimitsOfSizeOfLeftOp (F : C ⥤ Dᵒᵖ) [PreservesLimitsOfSize.{w, w'} F.leftOp] : +lemma preservesColimitsOfSize_of_leftOp (F : C ⥤ Dᵒᵖ) [PreservesLimitsOfSize.{w, w'} F.leftOp] : PreservesColimitsOfSize.{w, w'} F where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeOfLeftOp _ _ + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_of_leftOp _ _ /-- If `F.rightOp : C ⥤ Dᵒᵖ` preserves limits, then `F : Cᵒᵖ ⥤ D` preserves colimits. -/ -def preservesColimitsOfSizeOfRightOp (F : Cᵒᵖ ⥤ D) [PreservesLimitsOfSize.{w, w'} F.rightOp] : +lemma preservesColimitsOfSize_of_rightOp (F : Cᵒᵖ ⥤ D) [PreservesLimitsOfSize.{w, w'} F.rightOp] : PreservesColimitsOfSize.{w, w'} F where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeOfRightOp _ _ + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_of_rightOp _ _ /-- If `F.unop : C ⥤ D` preserves limits, then `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits. -/ -def preservesColimitsOfSizeOfUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimitsOfSize.{w, w'} F.unop] : +lemma preservesColimitsOfSize_of_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimitsOfSize.{w, w'} F.unop] : PreservesColimitsOfSize.{w, w'} F where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeOfUnop _ _ + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_of_unop _ _ /-- If `F : C ⥤ D` preserves colimits, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits. -/ -def preservesLimitsOp (F : C ⥤ D) [PreservesColimits F] : PreservesLimits F.op where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeOp _ _ +lemma preservesLimits_op (F : C ⥤ D) [PreservesColimits F] : PreservesLimits F.op where + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_op _ _ /-- If `F : C ⥤ Dᵒᵖ` preserves colimits, then `F.leftOp : Cᵒᵖ ⥤ D` preserves limits. -/ -def preservesLimitsLeftOp (F : C ⥤ Dᵒᵖ) [PreservesColimits F] : PreservesLimits F.leftOp where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeLeftOp _ _ +lemma preservesLimits_leftOp (F : C ⥤ Dᵒᵖ) [PreservesColimits F] : PreservesLimits F.leftOp where + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_leftOp _ _ /-- If `F : Cᵒᵖ ⥤ D` preserves colimits, then `F.rightOp : C ⥤ Dᵒᵖ` preserves limits. -/ -def preservesLimitsRightOp (F : Cᵒᵖ ⥤ D) [PreservesColimits F] : PreservesLimits F.rightOp where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeRightOp _ _ +lemma preservesLimits_rightOp (F : Cᵒᵖ ⥤ D) [PreservesColimits F] : PreservesLimits F.rightOp where + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_rightOp _ _ /-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits, then `F.unop : C ⥤ D` preserves limits. -/ -def preservesLimitsUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimits F] : PreservesLimits F.unop where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeUnop _ _ +lemma preservesLimits_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimits F] : PreservesLimits F.unop where + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_unop _ _ /-- If `F : C ⥤ D` preserves limits, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits. -/ -def perservesColimitsOp (F : C ⥤ D) [PreservesLimits F] : PreservesColimits F.op where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeOp _ _ +lemma perservesColimits_op (F : C ⥤ D) [PreservesLimits F] : PreservesColimits F.op where + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_op _ _ /-- If `F : C ⥤ Dᵒᵖ` preserves limits, then `F.leftOp : Cᵒᵖ ⥤ D` preserves colimits. -/ -def preservesColimitsLeftOp (F : C ⥤ Dᵒᵖ) [PreservesLimits F] : PreservesColimits F.leftOp where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeLeftOp _ _ +lemma preservesColimits_leftOp (F : C ⥤ Dᵒᵖ) [PreservesLimits F] : PreservesColimits F.leftOp where + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_leftOp _ _ /-- If `F : Cᵒᵖ ⥤ D` preserves limits, then `F.rightOp : C ⥤ Dᵒᵖ` preserves colimits. -/ -def preservesColimitsRightOp (F : Cᵒᵖ ⥤ D) [PreservesLimits F] : PreservesColimits F.rightOp where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeRightOp _ _ +lemma preservesColimits_rightOp (F : Cᵒᵖ ⥤ D) [PreservesLimits F] : + PreservesColimits F.rightOp where + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_rightOp _ _ /-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits, then `F.unop : C ⥤ D` preserves colimits. -/ -def preservesColimitsUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimits F] : PreservesColimits F.unop where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeUnop _ _ +lemma preservesColimits_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimits F] : PreservesColimits F.unop where + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_unop _ _ /-- If `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits, then `F : C ⥤ D` preserves limits. -/ -def preservesLimitsOfOp (F : C ⥤ D) [PreservesColimits F.op] : PreservesLimits F where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeOfOp _ _ +lemma preservesLimits_of_op (F : C ⥤ D) [PreservesColimits F.op] : PreservesLimits F where + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_of_op _ _ /-- If `F.leftOp : Cᵒᵖ ⥤ D` preserves colimits, then `F : C ⥤ Dᵒᵖ` preserves limits. -/ -def preservesLimitsOfLeftOp (F : C ⥤ Dᵒᵖ) [PreservesColimits F.leftOp] : PreservesLimits F where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeOfLeftOp _ _ +lemma preservesLimits_of_leftOp (F : C ⥤ Dᵒᵖ) [PreservesColimits F.leftOp] : PreservesLimits F where + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_of_leftOp _ _ /-- If `F.rightOp : C ⥤ Dᵒᵖ` preserves colimits, then `F : Cᵒᵖ ⥤ D` preserves limits. -/ -def preservesLimitsOfRightOp (F : Cᵒᵖ ⥤ D) [PreservesColimits F.rightOp] : PreservesLimits F where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeOfRightOp _ _ +lemma preservesLimits_of_rightOp (F : Cᵒᵖ ⥤ D) [PreservesColimits F.rightOp] : + PreservesLimits F where + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_of_rightOp _ _ /-- If `F.unop : C ⥤ D` preserves colimits, then `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits. -/ -def preservesLimitsOfUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimits F.unop] : PreservesLimits F where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeOfUnop _ _ +lemma preservesLimits_of_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimits F.unop] : PreservesLimits F where + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_of_unop _ _ /-- If `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits, then `F : C ⥤ D` preserves colimits. -/ -def preservesColimitsOfOp (F : C ⥤ D) [PreservesLimits F.op] : PreservesColimits F where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeOfOp _ _ +lemma preservesColimits_of_op (F : C ⥤ D) [PreservesLimits F.op] : PreservesColimits F where + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_of_op _ _ /-- If `F.leftOp : Cᵒᵖ ⥤ D` preserves limits, then `F : C ⥤ Dᵒᵖ` preserves colimits. -/ -def preservesColimitsOfLeftOp (F : C ⥤ Dᵒᵖ) [PreservesLimits F.leftOp] : PreservesColimits F where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeOfLeftOp _ _ +lemma preservesColimits_of_leftOp (F : C ⥤ Dᵒᵖ) [PreservesLimits F.leftOp] : + PreservesColimits F where + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_of_leftOp _ _ /-- If `F.rightOp : C ⥤ Dᵒᵖ` preserves limits, then `F : Cᵒᵖ ⥤ D` preserves colimits. -/ -def preservesColimitsOfRightOp (F : Cᵒᵖ ⥤ D) [PreservesLimits F.rightOp] : PreservesColimits F where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeOfRightOp _ _ +lemma preservesColimits_of_rightOp (F : Cᵒᵖ ⥤ D) [PreservesLimits F.rightOp] : + PreservesColimits F where + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_of_rightOp _ _ /-- If `F.unop : C ⥤ D` preserves limits, then `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits. -/ -def preservesColimitsOfUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimits F.unop] : PreservesColimits F where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeOfUnop _ _ +lemma preservesColimits_of_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimits F.unop] : PreservesColimits F where + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_of_unop _ _ /-- If `F : C ⥤ D` preserves finite colimits, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite limits. -/ -def preservesFiniteLimitsOp (F : C ⥤ D) [PreservesFiniteColimits F] : +lemma preservesFiniteLimits_op (F : C ⥤ D) [PreservesFiniteColimits F] : PreservesFiniteLimits F.op where - preservesFiniteLimits J _ _ := preservesLimitsOfShapeOp J F + preservesFiniteLimits J _ _ := preservesLimitsOfShape_op J F /-- If `F : C ⥤ Dᵒᵖ` preserves finite colimits, then `F.leftOp : Cᵒᵖ ⥤ D` preserves finite limits. -/ -def preservesFiniteLimitsLeftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteColimits F] : +lemma preservesFiniteLimits_leftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteColimits F] : PreservesFiniteLimits F.leftOp where - preservesFiniteLimits J _ _ := preservesLimitsOfShapeLeftOp J F + preservesFiniteLimits J _ _ := preservesLimitsOfShape_leftOp J F /-- If `F : Cᵒᵖ ⥤ D` preserves finite colimits, then `F.rightOp : C ⥤ Dᵒᵖ` preserves finite limits. -/ -def preservesFiniteLimitsRightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteColimits F] : +lemma preservesFiniteLimits_rightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteColimits F] : PreservesFiniteLimits F.rightOp where - preservesFiniteLimits J _ _ := preservesLimitsOfShapeRightOp J F + preservesFiniteLimits J _ _ := preservesLimitsOfShape_rightOp J F /-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite colimits, then `F.unop : C ⥤ D` preserves finite limits. -/ -def preservesFiniteLimitsUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteColimits F] : +lemma preservesFiniteLimits_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteColimits F] : PreservesFiniteLimits F.unop where - preservesFiniteLimits J _ _ := preservesLimitsOfShapeUnop J F + preservesFiniteLimits J _ _ := preservesLimitsOfShape_unop J F /-- If `F : C ⥤ D` preserves finite limits, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite colimits. -/ -def preservesFiniteColimitsOp (F : C ⥤ D) [PreservesFiniteLimits F] : +lemma preservesFiniteColimits_op (F : C ⥤ D) [PreservesFiniteLimits F] : PreservesFiniteColimits F.op where - preservesFiniteColimits J _ _ := preservesColimitsOfShapeOp J F + preservesFiniteColimits J _ _ := preservesColimitsOfShape_op J F /-- If `F : C ⥤ Dᵒᵖ` preserves finite limits, then `F.leftOp : Cᵒᵖ ⥤ D` preserves finite colimits. -/ -def preservesFiniteColimitsLeftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteLimits F] : +lemma preservesFiniteColimits_leftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteLimits F] : PreservesFiniteColimits F.leftOp where - preservesFiniteColimits J _ _ := preservesColimitsOfShapeLeftOp J F + preservesFiniteColimits J _ _ := preservesColimitsOfShape_leftOp J F /-- If `F : Cᵒᵖ ⥤ D` preserves finite limits, then `F.rightOp : C ⥤ Dᵒᵖ` preserves finite colimits. -/ -def preservesFiniteColimitsRightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteLimits F] : +lemma preservesFiniteColimits_rightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteLimits F] : PreservesFiniteColimits F.rightOp where - preservesFiniteColimits J _ _ := preservesColimitsOfShapeRightOp J F + preservesFiniteColimits J _ _ := preservesColimitsOfShape_rightOp J F /-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite limits, then `F.unop : C ⥤ D` preserves finite colimits. -/ -def preservesFiniteColimitsUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteLimits F] : +lemma preservesFiniteColimits_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteLimits F] : PreservesFiniteColimits F.unop where - preservesFiniteColimits J _ _ := preservesColimitsOfShapeUnop J F + preservesFiniteColimits J _ _ := preservesColimitsOfShape_unop J F /-- If `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite colimits, then `F : C ⥤ D` preserves finite limits. -/ -def preservesFiniteLimitsOfOp (F : C ⥤ D) [PreservesFiniteColimits F.op] : +lemma preservesFiniteLimits_of_op (F : C ⥤ D) [PreservesFiniteColimits F.op] : PreservesFiniteLimits F where - preservesFiniteLimits J _ _ := preservesLimitsOfShapeOfOp J F + preservesFiniteLimits J _ _ := preservesLimitsOfShape_of_op J F /-- If `F.leftOp : Cᵒᵖ ⥤ D` preserves finite colimits, then `F : C ⥤ Dᵒᵖ` preserves finite limits. -/ -def preservesFiniteLimitsOfLeftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteColimits F.leftOp] : +lemma preservesFiniteLimits_of_leftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteColimits F.leftOp] : PreservesFiniteLimits F where - preservesFiniteLimits J _ _ := preservesLimitsOfShapeOfLeftOp J F + preservesFiniteLimits J _ _ := preservesLimitsOfShape_of_leftOp J F /-- If `F.rightOp : C ⥤ Dᵒᵖ` preserves finite colimits, then `F : Cᵒᵖ ⥤ D` preserves finite limits. -/ -def preservesFiniteLimitsOfRightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteColimits F.rightOp] : +lemma preservesFiniteLimits_of_rightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteColimits F.rightOp] : PreservesFiniteLimits F where - preservesFiniteLimits J _ _ := preservesLimitsOfShapeOfRightOp J F + preservesFiniteLimits J _ _ := preservesLimitsOfShape_of_rightOp J F /-- If `F.unop : C ⥤ D` preserves finite colimits, then `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite limits. -/ -def preservesFiniteLimitsOfUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteColimits F.unop] : +lemma preservesFiniteLimits_of_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteColimits F.unop] : PreservesFiniteLimits F where - preservesFiniteLimits J _ _ := preservesLimitsOfShapeOfUnop J F + preservesFiniteLimits J _ _ := preservesLimitsOfShape_of_unop J F /-- If `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite limits, then `F : C ⥤ D` preserves finite colimits. -/ -def preservesFiniteColimitsOfOp (F : C ⥤ D) [PreservesFiniteLimits F.op] : +lemma preservesFiniteColimits_of_op (F : C ⥤ D) [PreservesFiniteLimits F.op] : PreservesFiniteColimits F where - preservesFiniteColimits J _ _ := preservesColimitsOfShapeOfOp J F + preservesFiniteColimits J _ _ := preservesColimitsOfShape_of_op J F /-- If `F.leftOp : Cᵒᵖ ⥤ D` preserves finite limits, then `F : C ⥤ Dᵒᵖ` preserves finite colimits. -/ -def preservesFiniteColimitsOfLeftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteLimits F.leftOp] : +lemma preservesFiniteColimits_of_leftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteLimits F.leftOp] : PreservesFiniteColimits F where - preservesFiniteColimits J _ _ := preservesColimitsOfShapeOfLeftOp J F + preservesFiniteColimits J _ _ := preservesColimitsOfShape_of_leftOp J F /-- If `F.rightOp : C ⥤ Dᵒᵖ` preserves finite limits, then `F : Cᵒᵖ ⥤ D` preserves finite colimits. -/ -def preservesFiniteColimitsOfRightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteLimits F.rightOp] : +lemma preservesFiniteColimits_of_rightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteLimits F.rightOp] : PreservesFiniteColimits F where - preservesFiniteColimits J _ _ := preservesColimitsOfShapeOfRightOp J F + preservesFiniteColimits J _ _ := preservesColimitsOfShape_of_rightOp J F /-- If `F.unop : C ⥤ D` preserves finite limits, then `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite colimits. -/ -def preservesFiniteColimitsOfUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteLimits F.unop] : +lemma preservesFiniteColimits_of_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteLimits F.unop] : PreservesFiniteColimits F where - preservesFiniteColimits J _ _ := preservesColimitsOfShapeOfUnop J F - + preservesFiniteColimits J _ _ := preservesColimitsOfShape_of_unop J F /-- If `F : C ⥤ D` preserves finite coproducts, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite products. -/ -def preservesFiniteProductsOp (F : C ⥤ D) [PreservesFiniteCoproducts F] : +lemma preservesFiniteProducts_op (F : C ⥤ D) [PreservesFiniteCoproducts F] : PreservesFiniteProducts F.op where preserves J _ := by - apply (config := { allowSynthFailures := true }) preservesLimitsOfShapeOp - exact preservesColimitsOfShapeOfEquiv (Discrete.opposite J).symm _ + apply (config := { allowSynthFailures := true }) preservesLimitsOfShape_op + exact preservesColimitsOfShape_of_equiv (Discrete.opposite J).symm _ /-- If `F : C ⥤ Dᵒᵖ` preserves finite coproducts, then `F.leftOp : Cᵒᵖ ⥤ D` preserves finite products. -/ -def preservesFiniteProductsLeftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteCoproducts F] : +lemma preservesFiniteProducts_leftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteCoproducts F] : PreservesFiniteProducts F.leftOp where preserves J _ := by - apply (config := { allowSynthFailures := true }) preservesLimitsOfShapeLeftOp - exact preservesColimitsOfShapeOfEquiv (Discrete.opposite J).symm _ + apply (config := { allowSynthFailures := true }) preservesLimitsOfShape_leftOp + exact preservesColimitsOfShape_of_equiv (Discrete.opposite J).symm _ /-- If `F : Cᵒᵖ ⥤ D` preserves finite coproducts, then `F.rightOp : C ⥤ Dᵒᵖ` preserves finite products. -/ -def preservesFiniteProductsRightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteCoproducts F] : +lemma preservesFiniteProducts_rightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteCoproducts F] : PreservesFiniteProducts F.rightOp where preserves J _ := by - apply (config := { allowSynthFailures := true }) preservesLimitsOfShapeRightOp - exact preservesColimitsOfShapeOfEquiv (Discrete.opposite J).symm _ + apply (config := { allowSynthFailures := true }) preservesLimitsOfShape_rightOp + exact preservesColimitsOfShape_of_equiv (Discrete.opposite J).symm _ /-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite coproducts, then `F.unop : C ⥤ D` preserves finite products. -/ -def preservesFiniteProductsUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteCoproducts F] : +lemma preservesFiniteProducts_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteCoproducts F] : PreservesFiniteProducts F.unop where preserves J _ := by - apply (config := { allowSynthFailures := true }) preservesLimitsOfShapeUnop - exact preservesColimitsOfShapeOfEquiv (Discrete.opposite J).symm _ + apply (config := { allowSynthFailures := true }) preservesLimitsOfShape_unop + exact preservesColimitsOfShape_of_equiv (Discrete.opposite J).symm _ /-- If `F : C ⥤ D` preserves finite products, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite coproducts. -/ -def preservesFiniteCoproductsOp (F : C ⥤ D) [PreservesFiniteProducts F] : +lemma preservesFiniteCoproducts_op (F : C ⥤ D) [PreservesFiniteProducts F] : PreservesFiniteCoproducts F.op where preserves J _ := by - apply (config := { allowSynthFailures := true }) preservesColimitsOfShapeOp - exact preservesLimitsOfShapeOfEquiv (Discrete.opposite J).symm _ + apply (config := { allowSynthFailures := true }) preservesColimitsOfShape_op + exact preservesLimitsOfShape_of_equiv (Discrete.opposite J).symm _ /-- If `F : C ⥤ Dᵒᵖ` preserves finite products, then `F.leftOp : Cᵒᵖ ⥤ D` preserves finite coproducts. -/ -def preservesFiniteCoproductsLeftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteProducts F] : +lemma preservesFiniteCoproducts_leftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteProducts F] : PreservesFiniteCoproducts F.leftOp where preserves J _ := by - apply (config := { allowSynthFailures := true }) preservesColimitsOfShapeLeftOp - exact preservesLimitsOfShapeOfEquiv (Discrete.opposite J).symm _ + apply (config := { allowSynthFailures := true }) preservesColimitsOfShape_leftOp + exact preservesLimitsOfShape_of_equiv (Discrete.opposite J).symm _ /-- If `F : Cᵒᵖ ⥤ D` preserves finite products, then `F.rightOp : C ⥤ Dᵒᵖ` preserves finite coproducts. -/ -def preservesFiniteCoproductsRightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteProducts F] : +lemma preservesFiniteCoproducts_rightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteProducts F] : PreservesFiniteCoproducts F.rightOp where preserves J _ := by - apply (config := { allowSynthFailures := true }) preservesColimitsOfShapeRightOp - exact preservesLimitsOfShapeOfEquiv (Discrete.opposite J).symm _ + apply (config := { allowSynthFailures := true }) preservesColimitsOfShape_rightOp + exact preservesLimitsOfShape_of_equiv (Discrete.opposite J).symm _ /-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite products, then `F.unop : C ⥤ D` preserves finite coproducts. -/ -def preservesFiniteCoproductsUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteProducts F] : +lemma preservesFiniteCoproducts_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteProducts F] : PreservesFiniteCoproducts F.unop where preserves J _ := by - apply (config := { allowSynthFailures := true }) preservesColimitsOfShapeUnop - exact preservesLimitsOfShapeOfEquiv (Discrete.opposite J).symm _ + apply (config := { allowSynthFailures := true }) preservesColimitsOfShape_unop + exact preservesLimitsOfShape_of_equiv (Discrete.opposite J).symm _ end CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Presheaf.lean b/Mathlib/CategoryTheory/Limits/Preserves/Presheaf.lean index f2a3c01bd9fc0..7750b104f48ec 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Presheaf.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Presheaf.lean @@ -172,9 +172,9 @@ attribute [local instance] PreservesFiniteLimitsOfIsFilteredCostructuredArrowYon One direction of Proposition 3.3.13 of [Kashiwara2006]. -/ -noncomputable def preservesFiniteLimitsOfIsFilteredCostructuredArrowYoneda +lemma preservesFiniteLimits_of_isFiltered_costructuredArrow_yoneda [IsFiltered (CostructuredArrow yoneda A)] : PreservesFiniteLimits A where - preservesFiniteLimits _ _ _ := ⟨fun {_} => preservesLimitOfIsIsoPost _ _⟩ + preservesFiniteLimits _ _ _ := ⟨fun {_} => preservesLimit_of_isIso_post _ _⟩ /-- If `C` is a small finitely cocomplete category and `A : Cᵒᵖ ⥤ Type u` is a presheaf, then `CostructuredArrow yoneda A` is filtered if and only if `A` preserves finite limits. @@ -183,7 +183,7 @@ Proposition 3.3.13 of [Kashiwara2006]. -/ theorem isFiltered_costructuredArrow_yoneda_iff_nonempty_preservesFiniteLimits : IsFiltered (CostructuredArrow yoneda A) ↔ Nonempty (PreservesFiniteLimits A) := - ⟨fun _ => ⟨preservesFiniteLimitsOfIsFilteredCostructuredArrowYoneda A⟩, + ⟨fun _ => ⟨preservesFiniteLimits_of_isFiltered_costructuredArrow_yoneda A⟩, fun ⟨_⟩ => isFiltered_costructuredArrow_yoneda_of_preservesFiniteLimits A⟩ end CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/BinaryProducts.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/BinaryProducts.lean index b72dd4a247aae..a9c9bfb413d2c 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/BinaryProducts.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/BinaryProducts.lean @@ -47,12 +47,12 @@ def isLimitMapConeBinaryFanEquiv : /-- The property of preserving products expressed in terms of binary fans. -/ def mapIsLimitOfPreservesOfIsLimit [PreservesLimit (pair X Y) G] (l : IsLimit (BinaryFan.mk f g)) : IsLimit (BinaryFan.mk (G.map f) (G.map g)) := - isLimitMapConeBinaryFanEquiv G f g (PreservesLimit.preserves l) + isLimitMapConeBinaryFanEquiv G f g (isLimitOfPreserves G l) /-- The property of reflecting products expressed in terms of binary fans. -/ def isLimitOfReflectsOfMapIsLimit [ReflectsLimit (pair X Y) G] (l : IsLimit (BinaryFan.mk (G.map f) (G.map g))) : IsLimit (BinaryFan.mk f g) := - ReflectsLimit.reflects ((isLimitMapConeBinaryFanEquiv G f g).symm l) + isLimitOfReflects G ((isLimitMapConeBinaryFanEquiv G f g).symm l) variable (X Y) variable [HasBinaryProduct X Y] @@ -69,9 +69,9 @@ variable [HasBinaryProduct (G.obj X) (G.obj Y)] /-- If the product comparison map for `G` at `(X,Y)` is an isomorphism, then `G` preserves the pair of `(X,Y)`. -/ -def PreservesLimitPair.ofIsoProdComparison [i : IsIso (prodComparison G X Y)] : +lemma PreservesLimitPair.of_iso_prod_comparison [i : IsIso (prodComparison G X Y)] : PreservesLimit (pair X Y) G := by - apply preservesLimitOfPreservesLimitCone (prodIsProd X Y) + apply preservesLimit_of_preserves_limit_cone (prodIsProd X Y) apply (isLimitMapConeBinaryFanEquiv _ _ _).symm _ refine @IsLimit.ofPointIso _ _ _ _ _ _ _ (limit.isLimit (pair (G.obj X) (G.obj Y))) ?_ apply i @@ -125,12 +125,12 @@ def isColimitMapCoconeBinaryCofanEquiv : /-- The property of preserving coproducts expressed in terms of binary cofans. -/ def mapIsColimitOfPreservesOfIsColimit [PreservesColimit (pair X Y) G] (l : IsColimit (BinaryCofan.mk f g)) : IsColimit (BinaryCofan.mk (G.map f) (G.map g)) := - isColimitMapCoconeBinaryCofanEquiv G f g (PreservesColimit.preserves l) + isColimitMapCoconeBinaryCofanEquiv G f g (isColimitOfPreserves G l) /-- The property of reflecting coproducts expressed in terms of binary cofans. -/ def isColimitOfReflectsOfMapIsColimit [ReflectsColimit (pair X Y) G] (l : IsColimit (BinaryCofan.mk (G.map f) (G.map g))) : IsColimit (BinaryCofan.mk f g) := - ReflectsColimit.reflects ((isColimitMapCoconeBinaryCofanEquiv G f g).symm l) + isColimitOfReflects G ((isColimitMapCoconeBinaryCofanEquiv G f g).symm l) variable (X Y) variable [HasBinaryCoproduct X Y] @@ -148,9 +148,9 @@ variable [HasBinaryCoproduct (G.obj X) (G.obj Y)] /-- If the coproduct comparison map for `G` at `(X,Y)` is an isomorphism, then `G` preserves the pair of `(X,Y)`. -/ -def PreservesColimitPair.ofIsoCoprodComparison [i : IsIso (coprodComparison G X Y)] : +lemma PreservesColimitPair.of_iso_coprod_comparison [i : IsIso (coprodComparison G X Y)] : PreservesColimit (pair X Y) G := by - apply preservesColimitOfPreservesColimitCocone (coprodIsCoprod X Y) + apply preservesColimit_of_preserves_colimit_cocone (coprodIsCoprod X Y) apply (isColimitMapCoconeBinaryCofanEquiv _ _ _).symm _ refine @IsColimit.ofPointIso _ _ _ _ _ _ _ (colimit.isColimit (pair (G.obj X) (G.obj Y))) ?_ apply i diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Biproducts.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Biproducts.lean index c3d4d473c6c16..550f8c8bea18d 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Biproducts.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Biproducts.lean @@ -85,8 +85,8 @@ variable {J : Type w₁} {K : Type w₂} /-- A functor `F` preserves biproducts of `f` if `F` maps every bilimit bicone over `f` to a bilimit bicone over `F.obj ∘ f`. -/ -class PreservesBiproduct (f : J → C) (F : C ⥤ D) [PreservesZeroMorphisms F] where - preserves : ∀ {b : Bicone f}, b.IsBilimit → (F.mapBicone b).IsBilimit +class PreservesBiproduct (f : J → C) (F : C ⥤ D) [PreservesZeroMorphisms F] : Prop where + preserves : ∀ {b : Bicone f}, b.IsBilimit → Nonempty (F.mapBicone b).IsBilimit attribute [inherit_doc PreservesBiproduct] PreservesBiproduct.preserves @@ -94,13 +94,13 @@ attribute [inherit_doc PreservesBiproduct] PreservesBiproduct.preserves bilimit bicone over `F.obj ∘ f`. -/ def isBilimitOfPreserves {f : J → C} (F : C ⥤ D) [PreservesZeroMorphisms F] [PreservesBiproduct f F] {b : Bicone f} (hb : b.IsBilimit) : (F.mapBicone b).IsBilimit := - PreservesBiproduct.preserves hb + (PreservesBiproduct.preserves hb).some variable (J) /-- A functor `F` preserves biproducts of shape `J` if it preserves biproducts of `f` for every `f : J → C`. -/ -class PreservesBiproductsOfShape (F : C ⥤ D) [PreservesZeroMorphisms F] where +class PreservesBiproductsOfShape (F : C ⥤ D) [PreservesZeroMorphisms F] : Prop where preserves : ∀ {f : J → C}, PreservesBiproduct f F attribute [inherit_doc PreservesBiproductsOfShape] PreservesBiproductsOfShape.preserves @@ -111,7 +111,7 @@ end Bicone /-- A functor `F` preserves finite biproducts if it preserves biproducts of shape `J` whenever `J` is a fintype. -/ -class PreservesFiniteBiproducts (F : C ⥤ D) [PreservesZeroMorphisms F] where +class PreservesFiniteBiproducts (F : C ⥤ D) [PreservesZeroMorphisms F] : Prop where preserves : ∀ {J : Type} [Fintype J], PreservesBiproductsOfShape J F attribute [inherit_doc PreservesFiniteBiproducts] PreservesFiniteBiproducts.preserves @@ -121,7 +121,7 @@ attribute [instance 100] PreservesFiniteBiproducts.preserves /-- A functor `F` preserves biproducts if it preserves biproducts of any shape `J` of size `w`. The usual notion of preservation of biproducts is recovered by choosing `w` to be the universe of the morphisms of `C`. -/ -class PreservesBiproducts (F : C ⥤ D) [PreservesZeroMorphisms F] where +class PreservesBiproducts (F : C ⥤ D) [PreservesZeroMorphisms F] : Prop where preserves : ∀ {J : Type w₁}, PreservesBiproductsOfShape J F attribute [inherit_doc PreservesBiproducts] PreservesBiproducts.preserves @@ -130,22 +130,22 @@ attribute [instance 100] PreservesBiproducts.preserves /-- Preserving biproducts at a bigger universe level implies preserving biproducts at a smaller universe level. -/ -def preservesBiproductsShrink (F : C ⥤ D) [PreservesZeroMorphisms F] +lemma preservesBiproducts_shrink (F : C ⥤ D) [PreservesZeroMorphisms F] [PreservesBiproducts.{max w₁ w₂} F] : PreservesBiproducts.{w₁} F := ⟨fun {_} => ⟨fun {_} => ⟨fun {b} ib => - ((F.mapBicone b).whiskerIsBilimitIff _).toFun - (isBilimitOfPreserves F ((b.whiskerIsBilimitIff Equiv.ulift.{w₂}).invFun ib))⟩⟩⟩ + ⟨((F.mapBicone b).whiskerIsBilimitIff _).toFun + (isBilimitOfPreserves F ((b.whiskerIsBilimitIff Equiv.ulift.{w₂}).invFun ib))⟩⟩⟩⟩ instance (priority := 100) preservesFiniteBiproductsOfPreservesBiproducts (F : C ⥤ D) [PreservesZeroMorphisms F] [PreservesBiproducts.{w₁} F] : PreservesFiniteBiproducts F where - preserves {J} _ := by letI := preservesBiproductsShrink.{0} F; infer_instance + preserves {J} _ := by letI := preservesBiproducts_shrink.{0} F; infer_instance /-- A functor `F` preserves binary biproducts of `X` and `Y` if `F` maps every bilimit bicone over `X` and `Y` to a bilimit bicone over `F.obj X` and `F.obj Y`. -/ -class PreservesBinaryBiproduct (X Y : C) (F : C ⥤ D) [PreservesZeroMorphisms F] where - preserves : ∀ {b : BinaryBicone X Y}, b.IsBilimit → (F.mapBinaryBicone b).IsBilimit +class PreservesBinaryBiproduct (X Y : C) (F : C ⥤ D) [PreservesZeroMorphisms F] : Prop where + preserves : ∀ {b : BinaryBicone X Y}, b.IsBilimit → Nonempty ((F.mapBinaryBicone b).IsBilimit) attribute [inherit_doc PreservesBinaryBiproduct] PreservesBinaryBiproduct.preserves @@ -154,20 +154,21 @@ attribute [inherit_doc PreservesBinaryBiproduct] PreservesBinaryBiproduct.preser def isBinaryBilimitOfPreserves {X Y : C} (F : C ⥤ D) [PreservesZeroMorphisms F] [PreservesBinaryBiproduct X Y F] {b : BinaryBicone X Y} (hb : b.IsBilimit) : (F.mapBinaryBicone b).IsBilimit := - PreservesBinaryBiproduct.preserves hb + (PreservesBinaryBiproduct.preserves hb).some /-- A functor `F` preserves binary biproducts if it preserves the binary biproduct of `X` and `Y` for all `X` and `Y`. -/ -class PreservesBinaryBiproducts (F : C ⥤ D) [PreservesZeroMorphisms F] where +class PreservesBinaryBiproducts (F : C ⥤ D) [PreservesZeroMorphisms F] : Prop where preserves : ∀ {X Y : C}, PreservesBinaryBiproduct X Y F := by infer_instance attribute [inherit_doc PreservesBinaryBiproducts] PreservesBinaryBiproducts.preserves /-- A functor that preserves biproducts of a pair preserves binary biproducts. -/ -def preservesBinaryBiproductOfPreservesBiproduct (F : C ⥤ D) [PreservesZeroMorphisms F] (X Y : C) - [PreservesBiproduct (pairFunction X Y) F] : PreservesBinaryBiproduct X Y F where - preserves {b} hb := - { isLimit := +lemma preservesBinaryBiproduct_of_preservesBiproduct (F : C ⥤ D) + [PreservesZeroMorphisms F] (X Y : C) [PreservesBiproduct (pairFunction X Y) F] : + PreservesBinaryBiproduct X Y F where + preserves {b} hb := ⟨{ + isLimit := IsLimit.ofIsoLimit ((IsLimit.postcomposeHomEquiv (diagramIsoPair _) _).symm (isBilimitOfPreserves F (b.toBiconeIsBilimit.symm hb)).isLimit) <| @@ -178,12 +179,12 @@ def preservesBinaryBiproductOfPreservesBiproduct (F : C ⥤ D) [PreservesZeroMor ((IsColimit.precomposeInvEquiv (diagramIsoPair _) _).symm (isBilimitOfPreserves F (b.toBiconeIsBilimit.symm hb)).isColimit) <| Cocones.ext (Iso.refl _) fun j => by - rcases j with ⟨⟨⟩⟩ <;> simp } + rcases j with ⟨⟨⟩⟩ <;> simp }⟩ /-- A functor that preserves biproducts of a pair preserves binary biproducts. -/ -def preservesBinaryBiproductsOfPreservesBiproducts (F : C ⥤ D) [PreservesZeroMorphisms F] +lemma preservesBinaryBiproducts_of_preservesBiproducts (F : C ⥤ D) [PreservesZeroMorphisms F] [PreservesBiproductsOfShape WalkingPair F] : PreservesBinaryBiproducts F where - preserves {X} Y := preservesBinaryBiproductOfPreservesBiproduct F X Y + preserves {X} Y := preservesBinaryBiproduct_of_preservesBiproduct F X Y attribute [instance 100] PreservesBinaryBiproducts.preserves @@ -257,13 +258,13 @@ variable [PreservesZeroMorphisms F] [PreservesBiproduct f F] instance hasBiproduct_of_preserves : HasBiproduct (F.obj ∘ f) := HasBiproduct.mk { bicone := F.mapBicone (biproduct.bicone f) - isBilimit := PreservesBiproduct.preserves (biproduct.isBilimit _) } + isBilimit := isBilimitOfPreserves _ (biproduct.isBilimit _) } /-- If `F` preserves a biproduct, we get a definitionally nice isomorphism `F.obj (⨁ f) ≅ ⨁ (F.obj ∘ f)`. -/ @[simp] def mapBiproduct : F.obj (⨁ f) ≅ ⨁ F.obj ∘ f := - biproduct.uniqueUpToIso _ (PreservesBiproduct.preserves (biproduct.isBilimit _)) + biproduct.uniqueUpToIso _ (isBilimitOfPreserves _ (biproduct.isBilimit _)) theorem mapBiproduct_hom : haveI : HasBiproduct fun j => F.obj (f j) := hasBiproduct_of_preserves F f @@ -341,13 +342,13 @@ variable [PreservesZeroMorphisms F] [PreservesBinaryBiproduct X Y F] instance hasBinaryBiproduct_of_preserves : HasBinaryBiproduct (F.obj X) (F.obj Y) := HasBinaryBiproduct.mk { bicone := F.mapBinaryBicone (BinaryBiproduct.bicone X Y) - isBilimit := PreservesBinaryBiproduct.preserves (BinaryBiproduct.isBilimit _ _) } + isBilimit := isBinaryBilimitOfPreserves F (BinaryBiproduct.isBilimit _ _) } /-- If `F` preserves a binary biproduct, we get a definitionally nice isomorphism `F.obj (X ⊞ Y) ≅ F.obj X ⊞ F.obj Y`. -/ @[simp] def mapBiprod : F.obj (X ⊞ Y) ≅ F.obj X ⊞ F.obj Y := - biprod.uniqueUpToIso _ _ (PreservesBinaryBiproduct.preserves (BinaryBiproduct.isBilimit _ _)) + biprod.uniqueUpToIso _ _ (isBinaryBilimitOfPreserves F (BinaryBiproduct.isBilimit _ _)) theorem mapBiprod_hom : (mapBiprod F X Y).hom = biprod.lift (F.map biprod.fst) (F.map biprod.snd) := rfl diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Equalizers.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Equalizers.lean index 0eece82e3b677..3e3b781aeb0d8 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Equalizers.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Equalizers.lean @@ -46,13 +46,13 @@ def isLimitMapConeForkEquiv : /-- The property of preserving equalizers expressed in terms of forks. -/ def isLimitForkMapOfIsLimit [PreservesLimit (parallelPair f g) G] (l : IsLimit (Fork.ofι h w)) : IsLimit (Fork.ofι (G.map h) (by simp only [← G.map_comp, w]) : Fork (G.map f) (G.map g)) := - isLimitMapConeForkEquiv G w (PreservesLimit.preserves l) + isLimitMapConeForkEquiv G w (isLimitOfPreserves G l) /-- The property of reflecting equalizers expressed in terms of forks. -/ def isLimitOfIsLimitForkMap [ReflectsLimit (parallelPair f g) G] (l : IsLimit (Fork.ofι (G.map h) (by simp only [← G.map_comp, w]) : Fork (G.map f) (G.map g))) : IsLimit (Fork.ofι h w) := - ReflectsLimit.reflects ((isLimitMapConeForkEquiv G w).symm l) + isLimitOfReflects G ((isLimitMapConeForkEquiv G w).symm l) variable (f g) variable [HasEqualizer f g] @@ -72,12 +72,11 @@ variable [HasEqualizer (G.map f) (G.map g)] /-- If the equalizer comparison map for `G` at `(f,g)` is an isomorphism, then `G` preserves the equalizer of `(f,g)`. -/ -def PreservesEqualizer.ofIsoComparison [i : IsIso (equalizerComparison f g G)] : +lemma PreservesEqualizer.of_iso_comparison [i : IsIso (equalizerComparison f g G)] : PreservesLimit (parallelPair f g) G := by - apply preservesLimitOfPreservesLimitCone (equalizerIsEqualizer f g) + apply preservesLimit_of_preserves_limit_cone (equalizerIsEqualizer f g) apply (isLimitMapConeForkEquiv _ _).symm _ - refine @IsLimit.ofPointIso _ _ _ _ _ _ _ (limit.isLimit (parallelPair (G.map f) (G.map g))) ?_ - apply i + exact @IsLimit.ofPointIso _ _ _ _ _ _ _ (limit.isLimit (parallelPair (G.map f) (G.map g))) i variable [PreservesLimit (parallelPair f g) G] @@ -128,7 +127,7 @@ def isColimitCoforkMapOfIsColimit [PreservesColimit (parallelPair f g) G] (l : IsColimit (Cofork.ofπ h w)) : IsColimit (Cofork.ofπ (G.map h) (by simp only [← G.map_comp, w]) : Cofork (G.map f) (G.map g)) := - isColimitMapCoconeCoforkEquiv G w (PreservesColimit.preserves l) + isColimitMapCoconeCoforkEquiv G w (isColimitOfPreserves G l) /-- The property of reflecting coequalizers expressed in terms of coforks. -/ def isColimitOfIsColimitCoforkMap [ReflectsColimit (parallelPair f g) G] @@ -136,7 +135,7 @@ def isColimitOfIsColimitCoforkMap [ReflectsColimit (parallelPair f g) G] IsColimit (Cofork.ofπ (G.map h) (by simp only [← G.map_comp, w]) : Cofork (G.map f) (G.map g))) : IsColimit (Cofork.ofπ h w) := - ReflectsColimit.reflects ((isColimitMapCoconeCoforkEquiv G w).symm l) + isColimitOfReflects G ((isColimitMapCoconeCoforkEquiv G w).symm l) variable (f g) variable [HasCoequalizer f g] @@ -155,13 +154,12 @@ variable [HasCoequalizer (G.map f) (G.map g)] /-- If the coequalizer comparison map for `G` at `(f,g)` is an isomorphism, then `G` preserves the coequalizer of `(f,g)`. -/ -def ofIsoComparison [i : IsIso (coequalizerComparison f g G)] : +lemma of_iso_comparison [i : IsIso (coequalizerComparison f g G)] : PreservesColimit (parallelPair f g) G := by - apply preservesColimitOfPreservesColimitCocone (coequalizerIsCoequalizer f g) + apply preservesColimit_of_preserves_colimit_cocone (coequalizerIsCoequalizer f g) apply (isColimitMapCoconeCoforkEquiv _ _).symm _ - refine - @IsColimit.ofPointIso _ _ _ _ _ _ _ (colimit.isColimit (parallelPair (G.map f) (G.map g))) ?_ - apply i + exact + @IsColimit.ofPointIso _ _ _ _ _ _ _ (colimit.isColimit (parallelPair (G.map f) (G.map g))) i variable [PreservesColimit (parallelPair f g) G] @@ -227,7 +225,7 @@ theorem map_π_preserves_coequalizer_inv_colimMap_desc {X' Y' : D} (f' g' : X' instance (priority := 1) preservesSplitCoequalizers (f g : X ⟶ Y) [HasSplitCoequalizer f g] : PreservesColimit (parallelPair f g) G := by apply - preservesColimitOfPreservesColimitCocone + preservesColimit_of_preserves_colimit_cocone (HasSplitCoequalizer.isSplitCoequalizer f g).isCoequalizer apply (isColimitMapCoconeCoforkEquiv G _).symm @@ -236,7 +234,7 @@ instance (priority := 1) preservesSplitCoequalizers (f g : X ⟶ Y) [HasSplitCoe instance (priority := 1) preservesSplitEqualizers (f g : X ⟶ Y) [HasSplitEqualizer f g] : PreservesLimit (parallelPair f g) G := by apply - preservesLimitOfPreservesLimitCone + preservesLimit_of_preserves_limit_cone (HasSplitEqualizer.isSplitEqualizer f g).isEqualizer apply (isLimitMapConeForkEquiv G _).symm diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Kernels.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Kernels.lean index b6333a08c0d86..034c5e96b668a 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Kernels.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Kernels.lean @@ -91,7 +91,7 @@ def isLimitForkMapOfIsLimit' [PreservesLimit (parallelPair f 0) G] IsLimit (KernelFork.ofι (G.map h) (by simp only [← G.map_comp, w, Functor.map_zero]) : Fork (G.map f) 0) := - isLimitMapConeForkEquiv' G w (PreservesLimit.preserves l) + isLimitMapConeForkEquiv' G w (isLimitOfPreserves G l) variable (f) variable [HasKernel f] @@ -114,9 +114,9 @@ variable [HasKernel (G.map f)] /-- If the kernel comparison map for `G` at `f` is an isomorphism, then `G` preserves the kernel of `f`. -/ -def PreservesKernel.ofIsoComparison [i : IsIso (kernelComparison f G)] : +lemma PreservesKernel.of_iso_comparison [i : IsIso (kernelComparison f G)] : PreservesLimit (parallelPair f 0) G := by - apply preservesLimitOfPreservesLimitCone (kernelIsKernel f) + apply preservesLimit_of_preserves_limit_cone (kernelIsKernel f) apply (isLimitMapConeForkEquiv' G (kernel.condition f)).symm _ exact @IsLimit.ofPointIso _ _ _ _ _ _ _ (kernelIsKernel (G.map f)) i @@ -212,7 +212,7 @@ def isColimitCoforkMapOfIsColimit' [PreservesColimit (parallelPair f 0) G] IsColimit (CokernelCofork.ofπ (G.map h) (by simp only [← G.map_comp, w, Functor.map_zero]) : Cofork (G.map f) 0) := - isColimitMapCoconeCoforkEquiv' G w (PreservesColimit.preserves l) + isColimitMapCoconeCoforkEquiv' G w (isColimitOfPreserves G l) variable (f) variable [HasCokernel f] @@ -236,9 +236,9 @@ variable [HasCokernel (G.map f)] /-- If the cokernel comparison map for `G` at `f` is an isomorphism, then `G` preserves the cokernel of `f`. -/ -def PreservesCokernel.ofIsoComparison [i : IsIso (cokernelComparison f G)] : +lemma PreservesCokernel.of_iso_comparison [i : IsIso (cokernelComparison f G)] : PreservesColimit (parallelPair f 0) G := by - apply preservesColimitOfPreservesColimitCocone (cokernelIsCokernel f) + apply preservesColimit_of_preserves_colimit_cocone (cokernelIsCokernel f) apply (isColimitMapCoconeCoforkEquiv' G (cokernel.condition f)).symm _ exact @IsColimit.ofPointIso _ _ _ _ _ _ _ (cokernelIsCokernel (G.map f)) i @@ -275,32 +275,32 @@ end Cokernels variable (X Y : C) (G : C ⥤ D) [Functor.PreservesZeroMorphisms G] -noncomputable instance preservesKernelZero : +instance preservesKernel_zero : PreservesLimit (parallelPair (0 : X ⟶ Y) 0) G where - preserves {c} hc := by + preserves {c} hc := ⟨by have := KernelFork.IsLimit.isIso_ι c hc rfl refine (KernelFork.isLimitMapConeEquiv c G).symm ?_ refine IsLimit.ofIsoLimit (KernelFork.IsLimit.ofId _ (G.map_zero _ _)) ?_ - exact (Fork.ext (G.mapIso (asIso (Fork.ι c))).symm (by simp)) + exact (Fork.ext (G.mapIso (asIso (Fork.ι c))).symm (by simp))⟩ -noncomputable instance preservesCokernelZero : +noncomputable instance preservesCokernel_zero : PreservesColimit (parallelPair (0 : X ⟶ Y) 0) G where - preserves {c} hc := by + preserves {c} hc := ⟨by have := CokernelCofork.IsColimit.isIso_π c hc rfl refine (CokernelCofork.isColimitMapCoconeEquiv c G).symm ?_ refine IsColimit.ofIsoColimit (CokernelCofork.IsColimit.ofId _ (G.map_zero _ _)) ?_ - exact (Cofork.ext (G.mapIso (asIso (Cofork.π c))) (by simp)) + exact (Cofork.ext (G.mapIso (asIso (Cofork.π c))) (by simp))⟩ variable {X Y} /-- The kernel of a zero map is preserved by any functor which preserves zero morphisms. -/ -noncomputable def preservesKernelZero' (f : X ⟶ Y) (hf : f = 0) : +lemma preservesKernel_zero' (f : X ⟶ Y) (hf : f = 0) : PreservesLimit (parallelPair f 0) G := by rw [hf] infer_instance /-- The cokernel of a zero map is preserved by any functor which preserves zero morphisms. -/ -noncomputable def preservesCokernelZero' (f : X ⟶ Y) (hf : f = 0) : +lemma preservesCokernel_zero' (f : X ⟶ Y) (hf : f = 0) : PreservesColimit (parallelPair f 0) G := by rw [hf] infer_instance diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Products.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Products.lean index 9d1105f497bd6..6f15d54ba6cac 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Products.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Products.lean @@ -46,13 +46,13 @@ def isLimitMapConeFanMkEquiv {P : C} (g : ∀ j, P ⟶ f j) : def isLimitFanMkObjOfIsLimit [PreservesLimit (Discrete.functor f) G] {P : C} (g : ∀ j, P ⟶ f j) (t : IsLimit (Fan.mk _ g)) : IsLimit (Fan.mk (G.obj P) fun j => G.map (g j) : Fan fun j => G.obj (f j)) := - isLimitMapConeFanMkEquiv _ _ _ (PreservesLimit.preserves t) + isLimitMapConeFanMkEquiv _ _ _ (isLimitOfPreserves G t) /-- The property of reflecting products expressed in terms of fans. -/ def isLimitOfIsLimitFanMkObj [ReflectsLimit (Discrete.functor f) G] {P : C} (g : ∀ j, P ⟶ f j) (t : IsLimit (Fan.mk _ fun j => G.map (g j) : Fan fun j => G.obj (f j))) : IsLimit (Fan.mk P g) := - ReflectsLimit.reflects ((isLimitMapConeFanMkEquiv _ _ _).symm t) + isLimitOfReflects G ((isLimitMapConeFanMkEquiv _ _ _).symm t) section @@ -69,13 +69,12 @@ def isLimitOfHasProductOfPreservesLimit [PreservesLimit (Discrete.functor f) G] variable [HasProduct fun j : J => G.obj (f j)] /-- If `pi_comparison G f` is an isomorphism, then `G` preserves the limit of `f`. -/ -def PreservesProduct.ofIsoComparison [i : IsIso (piComparison G f)] : +lemma PreservesProduct.of_iso_comparison [i : IsIso (piComparison G f)] : PreservesLimit (Discrete.functor f) G := by - apply preservesLimitOfPreservesLimitCone (productIsProduct f) + apply preservesLimit_of_preserves_limit_cone (productIsProduct f) apply (isLimitMapConeFanMkEquiv _ _ _).symm _ - refine @IsLimit.ofPointIso _ _ _ _ _ _ _ - (limit.isLimit (Discrete.functor fun j : J => G.obj (f j))) ?_ - apply i + exact @IsLimit.ofPointIso _ _ _ _ _ _ _ + (limit.isLimit (Discrete.functor fun j : J => G.obj (f j))) i variable [PreservesLimit (Discrete.functor f) G] @@ -110,14 +109,14 @@ def isColimitMapCoconeCofanMkEquiv {P : C} (g : ∀ j, f j ⟶ P) : def isColimitCofanMkObjOfIsColimit [PreservesColimit (Discrete.functor f) G] {P : C} (g : ∀ j, f j ⟶ P) (t : IsColimit (Cofan.mk _ g)) : IsColimit (Cofan.mk (G.obj P) fun j => G.map (g j) : Cofan fun j => G.obj (f j)) := - isColimitMapCoconeCofanMkEquiv _ _ _ (PreservesColimit.preserves t) + isColimitMapCoconeCofanMkEquiv _ _ _ (isColimitOfPreserves G t) /-- The property of reflecting coproducts expressed in terms of cofans. -/ def isColimitOfIsColimitCofanMkObj [ReflectsColimit (Discrete.functor f) G] {P : C} (g : ∀ j, f j ⟶ P) (t : IsColimit (Cofan.mk _ fun j => G.map (g j) : Cofan fun j => G.obj (f j))) : IsColimit (Cofan.mk P g) := - ReflectsColimit.reflects ((isColimitMapCoconeCofanMkEquiv _ _ _).symm t) + isColimitOfReflects G ((isColimitMapCoconeCofanMkEquiv _ _ _).symm t) section @@ -133,13 +132,12 @@ def isColimitOfHasCoproductOfPreservesColimit [PreservesColimit (Discrete.functo variable [HasCoproduct fun j : J => G.obj (f j)] /-- If `sigma_comparison G f` is an isomorphism, then `G` preserves the colimit of `f`. -/ -def PreservesCoproduct.ofIsoComparison [i : IsIso (sigmaComparison G f)] : +lemma PreservesCoproduct.of_iso_comparison [i : IsIso (sigmaComparison G f)] : PreservesColimit (Discrete.functor f) G := by - apply preservesColimitOfPreservesColimitCocone (coproductIsCoproduct f) + apply preservesColimit_of_preserves_colimit_cocone (coproductIsCoproduct f) apply (isColimitMapCoconeCofanMkEquiv _ _ _).symm _ - refine @IsColimit.ofPointIso _ _ _ _ _ _ _ - (colimit.isColimit (Discrete.functor fun j : J => G.obj (f j))) ?_ - apply i + exact @IsColimit.ofPointIso _ _ _ _ _ _ _ + (colimit.isColimit (Discrete.functor fun j : J => G.obj (f j))) i variable [PreservesColimit (Discrete.functor f) G] diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Pullbacks.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Pullbacks.lean index d79fa92a4f909..0eb7d4f49fdba 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Pullbacks.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Pullbacks.lean @@ -80,13 +80,13 @@ def isLimitPullbackConeMapOfIsLimit [PreservesLimit (cospan f g) G] (l : IsLimit (PullbackCone.mk h k comm)) : have : G.map h ≫ G.map f = G.map k ≫ G.map g := by rw [← G.map_comp, ← G.map_comp,comm] IsLimit (PullbackCone.mk (G.map h) (G.map k) this) := - (PullbackCone.isLimitMapConeEquiv _ G).1 (PreservesLimit.preserves l) + (PullbackCone.isLimitMapConeEquiv _ G).1 (isLimitOfPreserves G l) /-- The property of reflecting pullbacks expressed in terms of binary fans. -/ def isLimitOfIsLimitPullbackConeMap [ReflectsLimit (cospan f g) G] (l : IsLimit (PullbackCone.mk (G.map h) (G.map k) (show G.map h ≫ G.map f = G.map k ≫ G.map g from by simp only [← G.map_comp,comm]))) : IsLimit (PullbackCone.mk h k comm) := - ReflectsLimit.reflects + isLimitOfReflects G ((PullbackCone.isLimitMapConeEquiv (PullbackCone.mk _ _ comm) G).2 l) variable (f g) [PreservesLimit (cospan f g) G] @@ -100,21 +100,21 @@ def isLimitOfHasPullbackOfPreservesLimit [HasPullback f g] : isLimitPullbackConeMapOfIsLimit G _ (pullbackIsPullback f g) /-- If `F` preserves the pullback of `f, g`, it also preserves the pullback of `g, f`. -/ -def preservesPullbackSymmetry : PreservesLimit (cospan g f) G where - preserves {c} hc := by +lemma preservesPullback_symmetry : PreservesLimit (cospan g f) G where + preserves {c} hc := ⟨by apply (IsLimit.postcomposeHomEquiv (diagramIsoCospan.{v₂} _) _).toFun apply IsLimit.ofIsoLimit _ (PullbackCone.isoMk _).symm apply PullbackCone.isLimitOfFlip apply (isLimitMapConePullbackConeEquiv _ _).toFun - · refine @PreservesLimit.preserves _ _ _ _ _ _ _ _ ?_ _ ?_ + · refine @isLimitOfPreserves _ _ _ _ _ _ _ _ _ ?_ ?_ + · apply PullbackCone.isLimitOfFlip + apply IsLimit.ofIsoLimit _ (PullbackCone.isoMk _) + exact (IsLimit.postcomposeHomEquiv (diagramIsoCospan.{v₁} _) _).invFun hc · dsimp infer_instance - apply PullbackCone.isLimitOfFlip - apply IsLimit.ofIsoLimit _ (PullbackCone.isoMk _) - exact (IsLimit.postcomposeHomEquiv (diagramIsoCospan.{v₁} _) _).invFun hc · exact (c.π.naturality WalkingCospan.Hom.inr).symm.trans - (c.π.naturality WalkingCospan.Hom.inl : _) + (c.π.naturality WalkingCospan.Hom.inl : _)⟩ theorem hasPullback_of_preservesPullback [HasPullback f g] : HasPullback (G.map f) (G.map g) := ⟨⟨⟨_, isLimitPullbackConeMapOfIsLimit G _ (pullbackIsPullback _ _)⟩⟩⟩ @@ -207,14 +207,14 @@ def isColimitPushoutCoconeMapOfIsColimit [PreservesColimit (span f g) G] (l : IsColimit (PushoutCocone.mk h k comm)) : IsColimit (PushoutCocone.mk (G.map h) (G.map k) (show G.map f ≫ G.map h = G.map g ≫ G.map k from by simp only [← G.map_comp,comm] )) := - isColimitMapCoconePushoutCoconeEquiv G comm (PreservesColimit.preserves l) + isColimitMapCoconePushoutCoconeEquiv G comm (isColimitOfPreserves G l) /-- The property of reflecting pushouts expressed in terms of binary cofans. -/ def isColimitOfIsColimitPushoutCoconeMap [ReflectsColimit (span f g) G] (l : IsColimit (PushoutCocone.mk (G.map h) (G.map k) (show G.map f ≫ G.map h = G.map g ≫ G.map k from by simp only [← G.map_comp,comm]))) : IsColimit (PushoutCocone.mk h k comm) := - ReflectsColimit.reflects ((isColimitMapCoconePushoutCoconeEquiv G comm).symm l) + isColimitOfReflects G ((isColimitMapCoconePushoutCoconeEquiv G comm).symm l) variable (f g) [PreservesColimit (span f g) G] @@ -227,16 +227,16 @@ def isColimitOfHasPushoutOfPreservesColimit [i : HasPushout f g] : isColimitPushoutCoconeMapOfIsColimit G _ (pushoutIsPushout f g) /-- If `F` preserves the pushout of `f, g`, it also preserves the pushout of `g, f`. -/ -def preservesPushoutSymmetry : PreservesColimit (span g f) G where - preserves {c} hc := by +lemma preservesPushout_symmetry : PreservesColimit (span g f) G where + preserves {c} hc := ⟨by apply (IsColimit.precomposeHomEquiv (diagramIsoSpan.{v₂} _).symm _).toFun apply IsColimit.ofIsoColimit _ (PushoutCocone.isoMk _).symm apply PushoutCocone.isColimitOfFlip apply (isColimitMapCoconePushoutCoconeEquiv _ _).toFun - · refine @PreservesColimit.preserves _ _ _ _ _ _ _ _ ?_ _ ?_ -- Porting note: more TC coddling - · dsimp - infer_instance + · refine @isColimitOfPreserves _ _ _ _ _ _ _ _ _ ?_ ?_ -- Porting note: more TC coddling · exact PushoutCocone.flipIsColimit hc + · dsimp + infer_instance⟩ theorem hasPushout_of_preservesPushout [HasPushout f g] : HasPushout (G.map f) (G.map g) := ⟨⟨⟨_, isColimitPushoutCoconeMapOfIsColimit G _ (pushoutIsPushout _ _)⟩⟩⟩ @@ -290,12 +290,11 @@ variable [HasPullback f g] [HasPullback (G.map f) (G.map g)] /-- If the pullback comparison map for `G` at `(f,g)` is an isomorphism, then `G` preserves the pullback of `(f,g)`. -/ -def PreservesPullback.ofIsoComparison [i : IsIso (pullbackComparison G f g)] : +lemma PreservesPullback.of_iso_comparison [i : IsIso (pullbackComparison G f g)] : PreservesLimit (cospan f g) G := by - apply preservesLimitOfPreservesLimitCone (pullbackIsPullback f g) + apply preservesLimit_of_preserves_limit_cone (pullbackIsPullback f g) apply (isLimitMapConePullbackConeEquiv _ _).symm _ - refine @IsLimit.ofPointIso _ _ _ _ _ _ _ (limit.isLimit (cospan (G.map f) (G.map g))) ?_ - apply i + exact @IsLimit.ofPointIso _ _ _ _ _ _ _ (limit.isLimit (cospan (G.map f) (G.map g))) i variable [PreservesLimit (cospan f g) G] @@ -312,13 +311,12 @@ variable [HasPushout f g] [HasPushout (G.map f) (G.map g)] /-- If the pushout comparison map for `G` at `(f,g)` is an isomorphism, then `G` preserves the pushout of `(f,g)`. -/ -def PreservesPushout.ofIsoComparison [i : IsIso (pushoutComparison G f g)] : +lemma PreservesPushout.of_iso_comparison [i : IsIso (pushoutComparison G f g)] : PreservesColimit (span f g) G := by - apply preservesColimitOfPreservesColimitCocone (pushoutIsPushout f g) + apply preservesColimit_of_preserves_colimit_cocone (pushoutIsPushout f g) apply (isColimitMapCoconePushoutCoconeEquiv _ _).symm _ -- Porting note: apply no longer creates goals for instances - refine @IsColimit.ofPointIso _ _ _ _ _ _ _ (colimit.isColimit (span (G.map f) (G.map g))) ?_ - apply i + exact @IsColimit.ofPointIso _ _ _ _ _ _ _ (colimit.isColimit (span (G.map f) (G.map g))) i variable [PreservesColimit (span f g) G] diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Terminal.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Terminal.lean index 74e1a1309fbcc..8d97f5f95983f 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Terminal.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Terminal.lean @@ -42,12 +42,12 @@ def isLimitMapConeEmptyConeEquiv : /-- The property of preserving terminal objects expressed in terms of `IsTerminal`. -/ def IsTerminal.isTerminalObj [PreservesLimit (Functor.empty.{0} C) G] (l : IsTerminal X) : IsTerminal (G.obj X) := - isLimitMapConeEmptyConeEquiv G X (PreservesLimit.preserves l) + isLimitMapConeEmptyConeEquiv G X (isLimitOfPreserves G l) /-- The property of reflecting terminal objects expressed in terms of `IsTerminal`. -/ def IsTerminal.isTerminalOfObj [ReflectsLimit (Functor.empty.{0} C) G] (l : IsTerminal (G.obj X)) : IsTerminal X := - ReflectsLimit.reflects ((isLimitMapConeEmptyConeEquiv G X).symm l) + isLimitOfReflects G ((isLimitMapConeEmptyConeEquiv G X).symm l) /-- A functor that preserves and reflects terminal objects induces an equivalence on `IsTerminal`. -/ @@ -60,10 +60,9 @@ def IsTerminal.isTerminalIffObj [PreservesLimit (Functor.empty.{0} C) G] right_inv := by aesop_cat /-- Preserving the terminal object implies preserving all limits of the empty diagram. -/ -def preservesLimitsOfShapePemptyOfPreservesTerminal [PreservesLimit (Functor.empty.{0} C) G] : - PreservesLimitsOfShape (Discrete PEmpty) G where - preservesLimit := - preservesLimitOfIsoDiagram G (Functor.emptyExt (Functor.empty.{0} C) _) +lemma preservesLimitsOfShape_pempty_of_preservesTerminal [PreservesLimit (Functor.empty.{0} C) G] : + PreservesLimitsOfShape (Discrete PEmpty.{1}) G where + preservesLimit := preservesLimit_of_iso_diagram G (Functor.emptyExt (Functor.empty.{0} C) _) variable [HasTerminal C] @@ -90,21 +89,21 @@ variable [HasTerminal D] /-- If the terminal comparison map for `G` is an isomorphism, then `G` preserves terminal objects. -/ -def PreservesTerminal.ofIsoComparison [i : IsIso (terminalComparison G)] : - PreservesLimit (Functor.empty C) G := by - apply preservesLimitOfPreservesLimitCone terminalIsTerminal +lemma PreservesTerminal.of_iso_comparison [i : IsIso (terminalComparison G)] : + PreservesLimit (Functor.empty.{0} C) G := by + apply preservesLimit_of_preserves_limit_cone terminalIsTerminal apply (isLimitMapConeEmptyConeEquiv _ _).symm _ exact @IsLimit.ofPointIso _ _ _ _ _ _ _ (limit.isLimit (Functor.empty.{0} D)) i /-- If there is any isomorphism `G.obj ⊤ ⟶ ⊤`, then `G` preserves terminal objects. -/ -def preservesTerminalOfIsIso (f : G.obj (⊤_ C) ⟶ ⊤_ D) [i : IsIso f] : - PreservesLimit (Functor.empty C) G := by +lemma preservesTerminal_of_isIso (f : G.obj (⊤_ C) ⟶ ⊤_ D) [i : IsIso f] : + PreservesLimit (Functor.empty.{0} C) G := by rw [Subsingleton.elim f (terminalComparison G)] at i - exact PreservesTerminal.ofIsoComparison G + exact PreservesTerminal.of_iso_comparison G /-- If there is any isomorphism `G.obj ⊤ ≅ ⊤`, then `G` preserves terminal objects. -/ -def preservesTerminalOfIso (f : G.obj (⊤_ C) ≅ ⊤_ D) : PreservesLimit (Functor.empty C) G := - preservesTerminalOfIsIso G f.hom +lemma preservesTerminal_of_iso (f : G.obj (⊤_ C) ≅ ⊤_ D) : PreservesLimit (Functor.empty.{0} C) G := + preservesTerminal_of_isIso G f.hom variable [PreservesLimit (Functor.empty.{0} C) G] @@ -134,12 +133,12 @@ def isColimitMapCoconeEmptyCoconeEquiv : /-- The property of preserving initial objects expressed in terms of `IsInitial`. -/ def IsInitial.isInitialObj [PreservesColimit (Functor.empty.{0} C) G] (l : IsInitial X) : IsInitial (G.obj X) := - isColimitMapCoconeEmptyCoconeEquiv G X (PreservesColimit.preserves l) + isColimitMapCoconeEmptyCoconeEquiv G X (isColimitOfPreserves G l) /-- The property of reflecting initial objects expressed in terms of `IsInitial`. -/ def IsInitial.isInitialOfObj [ReflectsColimit (Functor.empty.{0} C) G] (l : IsInitial (G.obj X)) : IsInitial X := - ReflectsColimit.reflects ((isColimitMapCoconeEmptyCoconeEquiv G X).symm l) + isColimitOfReflects G ((isColimitMapCoconeEmptyCoconeEquiv G X).symm l) /-- A functor that preserves and reflects initial objects induces an equivalence on `IsInitial`. -/ def IsInitial.isInitialIffObj [PreservesColimit (Functor.empty.{0} C) G] @@ -151,10 +150,11 @@ def IsInitial.isInitialIffObj [PreservesColimit (Functor.empty.{0} C) G] right_inv := by aesop_cat /-- Preserving the initial object implies preserving all colimits of the empty diagram. -/ -def preservesColimitsOfShapePemptyOfPreservesInitial [PreservesColimit (Functor.empty.{0} C) G] : - PreservesColimitsOfShape (Discrete PEmpty) G where +lemma preservesColimitsOfShape_pempty_of_preservesInitial + [PreservesColimit (Functor.empty.{0} C) G] : + PreservesColimitsOfShape (Discrete PEmpty.{1}) G where preservesColimit := - preservesColimitOfIsoDiagram G (Functor.emptyExt (Functor.empty.{0} C) _) + preservesColimit_of_iso_diagram G (Functor.emptyExt (Functor.empty.{0} C) _) variable [HasInitial C] @@ -181,21 +181,22 @@ variable [HasInitial D] /-- If the initial comparison map for `G` is an isomorphism, then `G` preserves initial objects. -/ -def PreservesInitial.ofIsoComparison [i : IsIso (initialComparison G)] : - PreservesColimit (Functor.empty C) G := by - apply preservesColimitOfPreservesColimitCocone initialIsInitial +lemma PreservesInitial.of_iso_comparison [i : IsIso (initialComparison G)] : + PreservesColimit (Functor.empty.{0} C) G := by + apply preservesColimit_of_preserves_colimit_cocone initialIsInitial apply (isColimitMapCoconeEmptyCoconeEquiv _ _).symm _ exact @IsColimit.ofPointIso _ _ _ _ _ _ _ (colimit.isColimit (Functor.empty.{0} D)) i /-- If there is any isomorphism `⊥ ⟶ G.obj ⊥`, then `G` preserves initial objects. -/ -def preservesInitialOfIsIso (f : ⊥_ D ⟶ G.obj (⊥_ C)) [i : IsIso f] : - PreservesColimit (Functor.empty C) G := by +lemma preservesInitial_of_isIso (f : ⊥_ D ⟶ G.obj (⊥_ C)) [i : IsIso f] : + PreservesColimit (Functor.empty.{0} C) G := by rw [Subsingleton.elim f (initialComparison G)] at i - exact PreservesInitial.ofIsoComparison G + exact PreservesInitial.of_iso_comparison G /-- If there is any isomorphism `⊥ ≅ G.obj ⊥`, then `G` preserves initial objects. -/ -def preservesInitialOfIso (f : ⊥_ D ≅ G.obj (⊥_ C)) : PreservesColimit (Functor.empty C) G := - preservesInitialOfIsIso G f.hom +lemma preservesInitial_of_iso (f : ⊥_ D ≅ G.obj (⊥_ C)) : + PreservesColimit (Functor.empty.{0} C) G := + preservesInitial_of_isIso G f.hom variable [PreservesColimit (Functor.empty.{0} C) G] diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Zero.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Zero.lean index 1b3b19cbe4153..9378d2cf0aea9 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Zero.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Zero.lean @@ -157,15 +157,15 @@ instance (priority := 100) preservesZeroMorphisms_of_preserves_terminal_object variable (F) /-- Preserving zero morphisms implies preserving terminal objects. -/ -def preservesTerminalObjectOfPreservesZeroMorphisms [PreservesZeroMorphisms F] : - PreservesLimit (Functor.empty C) F := - preservesTerminalOfIso F <| +lemma preservesTerminalObject_of_preservesZeroMorphisms [PreservesZeroMorphisms F] : + PreservesLimit (Functor.empty.{0} C) F := + preservesTerminal_of_iso F <| F.mapIso HasZeroObject.zeroIsoTerminal.symm ≪≫ mapZeroObject F ≪≫ HasZeroObject.zeroIsoTerminal /-- Preserving zero morphisms implies preserving terminal objects. -/ -def preservesInitialObjectOfPreservesZeroMorphisms [PreservesZeroMorphisms F] : - PreservesColimit (Functor.empty C) F := - preservesInitialOfIso F <| +lemma preservesInitialObject_of_preservesZeroMorphisms [PreservesZeroMorphisms F] : + PreservesColimit (Functor.empty.{0} C) F := + preservesInitial_of_iso F <| HasZeroObject.zeroIsoInitial.symm ≪≫ (mapZeroObject F).symm ≪≫ (F.mapIso HasZeroObject.zeroIsoInitial.symm).symm @@ -176,25 +176,27 @@ section variable [HasZeroObject D] [HasZeroMorphisms D] (G : C ⥤ D) (hG : IsZero G) (J : Type*) [Category J] +include hG + /-- A zero functor preserves limits. -/ -def preservesLimitsOfShapeOfIsZero : PreservesLimitsOfShape J G where - preservesLimit {K} := ⟨fun _ => by +lemma preservesLimitsOfShape_of_isZero : PreservesLimitsOfShape J G where + preservesLimit {K} := ⟨fun _ => ⟨by rw [Functor.isZero_iff] at hG - exact IsLimit.ofIsZero _ ((K ⋙ G).isZero (fun X ↦ hG _)) (hG _)⟩ + exact IsLimit.ofIsZero _ ((K ⋙ G).isZero (fun X ↦ hG _)) (hG _)⟩⟩ /-- A zero functor preserves colimits. -/ -def preservesColimitsOfShapeOfIsZero : PreservesColimitsOfShape J G where - preservesColimit {K} := ⟨fun _ => by +lemma preservesColimitsOfShape_of_isZero : PreservesColimitsOfShape J G where + preservesColimit {K} := ⟨fun _ => ⟨by rw [Functor.isZero_iff] at hG - exact IsColimit.ofIsZero _ ((K ⋙ G).isZero (fun X ↦ hG _)) (hG _)⟩ + exact IsColimit.ofIsZero _ ((K ⋙ G).isZero (fun X ↦ hG _)) (hG _)⟩⟩ /-- A zero functor preserves limits. -/ -def preservesLimitsOfSizeOfIsZero : PreservesLimitsOfSize.{v, u} G where - preservesLimitsOfShape := G.preservesLimitsOfShapeOfIsZero hG _ +lemma preservesLimitsOfSize_of_isZero : PreservesLimitsOfSize.{v, u} G where + preservesLimitsOfShape := G.preservesLimitsOfShape_of_isZero hG _ /-- A zero functor preserves colimits. -/ -def preservesColimitsOfSizeOfIsZero : PreservesColimitsOfSize.{v, u} G where - preservesColimitsOfShape := G.preservesColimitsOfShapeOfIsZero hG _ +lemma preservesColimitsOfSize_of_isZero : PreservesColimitsOfSize.{v, u} G where + preservesColimitsOfShape := G.preservesColimitsOfShape_of_isZero hG _ end diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean b/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean index e11c09ddc138d..eb1da8048a51a 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean @@ -42,7 +42,6 @@ instance : PreservesLimitsOfSize.{w', w} uliftFunctor.{v, u} where preservesLimitsOfShape {J} := { preservesLimit := fun {K} => { preserves := fun {c} hc => by - apply Nonempty.some rw [Types.isLimit_iff ((uliftFunctor.{v, u}).mapCone c)] intro s hs obtain ⟨x, hx₁, hx₂⟩ := (Types.isLimit_iff c).mp ⟨hc⟩ _ ((sectionsEquiv K).symm ⟨s, hs⟩).2 @@ -138,11 +137,11 @@ The functor `uliftFunctor : Type u ⥤ Type (max u v)` preserves colimits of arb noncomputable instance : PreservesColimitsOfSize.{w', w} uliftFunctor.{v, u} where preservesColimitsOfShape {J _} := { preservesColimit := fun {F} ↦ - { preserves := fun {c} hc ↦ - { desc := fun lc x ↦ descFun hc lc x.down + { preserves := fun {c} hc ↦ ⟨{ + desc := fun lc x ↦ descFun hc lc x.down fac := fun lc j ↦ by ext ⟨⟩; apply congr_fun ((descFun_spec hc lc _).mp rfl j) uniq := fun lc f hf ↦ by ext ⟨⟩; apply congr_fun ((descFun_spec hc lc (f ∘ ULift.up)).mpr - fun j ↦ funext fun y ↦ congr_fun (hf j) ⟨y⟩) } } } + fun j ↦ funext fun y ↦ congr_fun (hf j) ⟨y⟩) }⟩ } } /-- The functor `uliftFunctor : Type u ⥤ Type (max u v)` creates `u`-small colimits. diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean b/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean index dabe16ad73949..7861c1e682600 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean @@ -70,7 +70,7 @@ theorem yonedaYonedaColimit_app_inv {X : C} : ((yonedaYonedaColimit F).app (op X noncomputable instance {X : C} : PreservesColimit F (coyoneda.obj (op (yoneda.obj X))) := by suffices IsIso (colimit.post F (coyoneda.obj (op (yoneda.obj X)))) from - preservesColimitOfIsIsoPost _ _ + preservesColimit_of_isIso_post _ _ suffices colimit.post F (coyoneda.obj (op (yoneda.obj X))) = (colimitObjIsoColimitCompEvaluation _ _).inv ≫ ((yonedaYonedaColimit F).app (op X)).inv from this ▸ inferInstance diff --git a/Mathlib/CategoryTheory/Limits/Presheaf.lean b/Mathlib/CategoryTheory/Limits/Presheaf.lean index ca7b9175b9037..9eca881c472cc 100644 --- a/Mathlib/CategoryTheory/Limits/Presheaf.lean +++ b/Mathlib/CategoryTheory/Limits/Presheaf.lean @@ -155,10 +155,11 @@ noncomputable def yonedaAdjunction : L ⊣ restrictedYoneda A := apply yonedaEquiv.injective simp [yonedaEquiv] } +include α in /-- Any left Kan extension along the Yoneda embedding preserves colimits. -/ -noncomputable def preservesColimitsOfSizeOfIsLeftKanExtension : +lemma preservesColimitsOfSize_of_isLeftKanExtension : PreservesColimitsOfSize.{v₃, u₃} L := - (yonedaAdjunction L α).leftAdjointPreservesColimits + (yonedaAdjunction L α).leftAdjoint_preservesColimits lemma isIso_of_isLeftKanExtension : IsIso α := (Functor.isPointwiseLeftKanExtensionOfIsLeftKanExtension _ α).isIso_hom @@ -166,9 +167,9 @@ lemma isIso_of_isLeftKanExtension : IsIso α := variable (A) /-- See Property 2 of https://ncatlab.org/nlab/show/Yoneda+extension#properties. -/ -noncomputable instance preservesColimitsOfSizeLeftKanExtension : +noncomputable instance preservesColimitsOfSize_leftKanExtension : PreservesColimitsOfSize.{v₃, u₃} (yoneda.leftKanExtension A) := - (yonedaAdjunction _ (yoneda.leftKanExtensionUnit A)).leftAdjointPreservesColimits + (yonedaAdjunction _ (yoneda.leftKanExtensionUnit A)).leftAdjoint_preservesColimits instance : IsIso (yoneda.leftKanExtensionUnit A) := isIso_of_isLeftKanExtension _ (yoneda.leftKanExtensionUnit A) @@ -257,12 +258,12 @@ instance [L.IsLeftKanExtension α] : IsIso α := lemma isLeftKanExtension_along_yoneda_iff : L.IsLeftKanExtension α ↔ - (IsIso α ∧ Nonempty (PreservesColimitsOfSize.{v₁, max u₁ v₁} L)) := by + (IsIso α ∧ PreservesColimitsOfSize.{v₁, max u₁ v₁} L) := by constructor · intro - exact ⟨inferInstance, ⟨preservesColimitsOfNatIso - (Functor.leftKanExtensionUnique _ (yoneda.leftKanExtensionUnit A) _ α)⟩⟩ - · rintro ⟨_, ⟨_⟩⟩ + exact ⟨inferInstance, preservesColimits_of_natIso + (Functor.leftKanExtensionUnique _ (yoneda.leftKanExtensionUnit A) _ α)⟩ + · rintro ⟨_, _⟩ apply Functor.LeftExtension.IsPointwiseLeftKanExtension.isLeftKanExtension (E := Functor.LeftExtension.mk _ α) intro P @@ -571,7 +572,7 @@ theorem final_toCostructuredArrow_comp_pre {c : Cocone (F ⋙ yoneda)} (hc : IsC (CostructuredArrow.toOverCompOverEquivPresheafCostructuredArrow c.pt).isoCompInverse apply IsTerminal.ofIso Over.mkIdTerminal - let isc : IsColimit ((Over.forget _).mapCocone _) := PreservesColimit.preserves + let isc : IsColimit ((Over.forget _).mapCocone _) := isColimitOfPreserves _ (colimit.isColimit ((c.toCostructuredArrow ⋙ CostructuredArrow.pre F yoneda c.pt) ⋙ CostructuredArrow.toOver yoneda c.pt)) exact Over.isoMk (hc.coconePointUniqueUpToIso isc) (hc.hom_ext fun i => by simp) diff --git a/Mathlib/CategoryTheory/Limits/Shapes/NormalMono/Basic.lean b/Mathlib/CategoryTheory/Limits/Shapes/NormalMono/Basic.lean index e16617bd21fec..8fba6cab9585a 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/NormalMono/Basic.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/NormalMono/Basic.lean @@ -59,16 +59,11 @@ def equivalenceReflectsNormalMono {D : Type u₂} [Category.{v₁} D] [HasZeroMo have reassoc' {W : D} (h : hf.Z ⟶ W) : F.map f ≫ hf.g ≫ h = 0 ≫ h := by rw [← Category.assoc, eq_whisker hf.w] simp [reassoc'] - isLimit := - @ReflectsLimit.reflects C _ D _ _ _ _ F _ _ <| - IsLimit.ofConeEquiv (Cones.postcomposeEquivalence (@compNatIso C _ _ _ _ _ D _ _ F _)) <| - IsLimit.ofIsoLimit - (IsLimit.ofIsoLimit - (IsKernel.ofCompIso _ _ (F.objObjPreimageIso hf.Z) (by - simp only [Functor.map_preimage, Category.assoc, Iso.inv_hom_id, Category.comp_id]) - hf.isLimit) - (ofιCongr (Category.comp_id _).symm)) - <| by apply Iso.symm; apply isoOfι -- Porting note: very fiddly unification here + isLimit := isLimitOfReflects F <| + IsLimit.ofConeEquiv (Cones.postcomposeEquivalence (compNatIso F)) <| + (IsLimit.ofIsoLimit (IsKernel.ofCompIso _ _ (F.objObjPreimageIso hf.Z) (by + simp only [Functor.map_preimage, Category.assoc, Iso.inv_hom_id, Category.comp_id]) + hf.isLimit)) (Fork.ext (Iso.refl _) (by simp [compNatIso, Fork.ι])) end @@ -160,14 +155,11 @@ def equivalenceReflectsNormalEpi {D : Type u₂} [Category.{v₁} D] [HasZeroMor W := F.objPreimage hf.W g := F.preimage ((F.objObjPreimageIso hf.W).hom ≫ hf.g) w := F.map_injective <| by simp [hf.w] - isColimit := - ReflectsColimit.reflects <| - IsColimit.ofCoconeEquiv (Cocones.precomposeEquivalence (compNatIso F).symm) <| - IsColimit.ofIsoColimit - (IsColimit.ofIsoColimit - (IsCokernel.ofIsoComp _ _ (F.objObjPreimageIso hf.W).symm (by simp) hf.isColimit) - (ofπCongr (Category.id_comp _).symm)) - <| by apply Iso.symm; apply isoOfπ + isColimit := isColimitOfReflects F <| + IsColimit.ofCoconeEquiv (Cocones.precomposeEquivalence (compNatIso F).symm) <| + (IsColimit.ofIsoColimit + (IsCokernel.ofIsoComp _ _ (F.objObjPreimageIso hf.W).symm (by simp) hf.isColimit) + (Cofork.ext (Iso.refl _) (by simp [compNatIso, Cofork.π]))) end diff --git a/Mathlib/CategoryTheory/Limits/VanKampen.lean b/Mathlib/CategoryTheory/Limits/VanKampen.lean index 429dea9a1687b..75e50cc0928fc 100644 --- a/Mathlib/CategoryTheory/Limits/VanKampen.lean +++ b/Mathlib/CategoryTheory/Limits/VanKampen.lean @@ -188,7 +188,7 @@ theorem IsUniversalColimit.of_mapCocone (G : C ⥤ D) {F : J ⥤ C} {c : Cocone [PreservesLimitsOfShape WalkingCospan G] [ReflectsColimitsOfShape J G] (hc : IsUniversalColimit (G.mapCocone c)) : IsUniversalColimit c := fun F' c' α f h hα H ↦ - ⟨ReflectsColimit.reflects (hc (G.mapCocone c') (whiskerRight α G) (G.map f) + ⟨isColimitOfReflects _ (hc (G.mapCocone c') (whiskerRight α G) (G.map f) (by ext j; simpa using G.congr_map (NatTrans.congr_app h j)) (hα.whiskerRight G) (fun j ↦ (H j).map G)).some⟩ @@ -282,7 +282,7 @@ theorem isVanKampenColimit_of_evaluation [HasPullbacks D] [HasColimitsOfShape J · rintro ⟨hc'⟩ j refine ⟨⟨(NatTrans.congr_app e j).symm⟩, ⟨evaluationJointlyReflectsLimits _ ?_⟩⟩ refine fun x => (isLimitMapConePullbackConeEquiv _ _).symm ?_ - exact ((this x).mp ⟨PreservesColimit.preserves hc'⟩ _).isLimit + exact ((this x).mp ⟨isColimitOfPreserves _ hc'⟩ _).isLimit · exact fun H => ⟨evaluationJointlyReflectsColimits _ fun x => ((this x).mpr fun j => (H j).map ((evaluation C D).obj x)).some⟩ @@ -297,8 +297,8 @@ theorem IsUniversalColimit.map_reflective [∀ X (f : X ⟶ Gl.obj c.pt), HasPullback (Gr.map f) (adj.unit.app c.pt)] [∀ X (f : X ⟶ Gl.obj c.pt), PreservesLimit (cospan (Gr.map f) (adj.unit.app c.pt)) Gl] : IsUniversalColimit (Gl.mapCocone c) := by - have := adj.rightAdjointPreservesLimits - have : PreservesColimitsOfSize.{u', v'} Gl := adj.leftAdjointPreservesColimits + have := adj.rightAdjoint_preservesLimits + have : PreservesColimitsOfSize.{u', v'} Gl := adj.leftAdjoint_preservesColimits intros F' c' α f h hα hc' have : HasPullback (Gl.map (Gr.map f)) (Gl.map (adj.unit.app c.pt)) := ⟨⟨_, isLimitPullbackConeMapOfIsLimit _ pullback.condition @@ -401,8 +401,8 @@ theorem IsVanKampenColimit.map_reflective [HasColimitsOfShape J C] [∀ X (f : X ⟶ Gl.obj c.pt), PreservesLimit (cospan (Gr.map f) (adj.unit.app c.pt)) Gl] [∀ X i (f : X ⟶ c.pt), PreservesLimit (cospan f (c.ι.app i)) Gl] : IsVanKampenColimit (Gl.mapCocone c) := by - have := adj.rightAdjointPreservesLimits - have : PreservesColimitsOfSize.{u', v'} Gl := adj.leftAdjointPreservesColimits + have := adj.rightAdjoint_preservesLimits + have : PreservesColimitsOfSize.{u', v'} Gl := adj.leftAdjoint_preservesColimits intro F' c' α f h hα refine ⟨?_, H.isUniversal.map_reflective adj c' α f h hα⟩ intro ⟨hc'⟩ j diff --git a/Mathlib/CategoryTheory/Limits/Yoneda.lean b/Mathlib/CategoryTheory/Limits/Yoneda.lean index ed9f8ff3d339e..fdb2d7ebc8f71 100644 --- a/Mathlib/CategoryTheory/Limits/Yoneda.lean +++ b/Mathlib/CategoryTheory/Limits/Yoneda.lean @@ -80,19 +80,19 @@ def Limits.coneOfSectionCompYoneda (F : J ⥤ Cᵒᵖ) (X : C) pt := Opposite.op X π := compYonedaSectionsEquiv F X s -noncomputable instance yonedaPreservesLimit (F : J ⥤ Cᵒᵖ) (X : C) : +instance yoneda_preservesLimit (F : J ⥤ Cᵒᵖ) (X : C) : PreservesLimit F (yoneda.obj X) where - preserves {c} hc := Nonempty.some (by + preserves {c} hc := by rw [Types.isLimit_iff] intro s hs exact ⟨(hc.lift (Limits.coneOfSectionCompYoneda F X ⟨s, hs⟩)).unop, fun j => Quiver.Hom.op_inj (hc.fac (Limits.coneOfSectionCompYoneda F X ⟨s, hs⟩) j), fun m hm => Quiver.Hom.op_inj (hc.uniq (Limits.coneOfSectionCompYoneda F X ⟨s, hs⟩) _ - (fun j => Quiver.Hom.unop_inj (hm j)))⟩) + (fun j => Quiver.Hom.unop_inj (hm j)))⟩ variable (J) in -noncomputable instance yonedaPreservesLimitsOfShape (X : C) : +noncomputable instance yoneda_preservesLimitsOfShape (X : C) : PreservesLimitsOfShape J (yoneda.obj X) where /-- The yoneda embeddings jointly reflect limits. -/ @@ -125,13 +125,13 @@ def Limits.coneOfSectionCompCoyoneda (F : J ⥤ C) (X : Cᵒᵖ) pt := X.unop π := compCoyonedaSectionsEquiv F X.unop s -noncomputable instance coyonedaPreservesLimit (F : J ⥤ C) (X : Cᵒᵖ) : +instance coyoneda_preservesLimit (F : J ⥤ C) (X : Cᵒᵖ) : PreservesLimit F (coyoneda.obj X) where - preserves {c} hc := Nonempty.some (by + preserves {c} hc := by rw [Types.isLimit_iff] intro s hs exact ⟨hc.lift (Limits.coneOfSectionCompCoyoneda F X ⟨s, hs⟩), hc.fac _, - hc.uniq (Limits.coneOfSectionCompCoyoneda F X ⟨s, hs⟩)⟩) + hc.uniq (Limits.coneOfSectionCompCoyoneda F X ⟨s, hs⟩)⟩ variable (J) in noncomputable instance coyonedaPreservesLimitsOfShape (X : Cᵒᵖ) : @@ -162,31 +162,31 @@ noncomputable def Limits.Cone.isLimitCoyonedaEquiv {F : J ⥤ C} (c : Cone F) : end /-- The yoneda embedding `yoneda.obj X : Cᵒᵖ ⥤ Type v` for `X : C` preserves limits. -/ -noncomputable instance yonedaPreservesLimits (X : C) : +instance yoneda_preservesLimits (X : C) : PreservesLimitsOfSize.{t, w} (yoneda.obj X) where /-- The coyoneda embedding `coyoneda.obj X : C ⥤ Type v` for `X : Cᵒᵖ` preserves limits. -/ -noncomputable instance coyonedaPreservesLimits (X : Cᵒᵖ) : +instance coyoneda_preservesLimits (X : Cᵒᵖ) : PreservesLimitsOfSize.{t, w} (coyoneda.obj X) where -noncomputable instance yonedaFunctorPreservesLimits : +instance yonedaFunctor_preservesLimits : PreservesLimitsOfSize.{t, w} (@yoneda C _) := by - apply preservesLimitsOfEvaluation + apply preservesLimits_of_evaluation intro K change PreservesLimitsOfSize (coyoneda.obj K) infer_instance -noncomputable instance coyonedaFunctorPreservesLimits : +noncomputable instance coyonedaFunctor_preservesLimits : PreservesLimitsOfSize.{t, w} (@coyoneda C _) := by - apply preservesLimitsOfEvaluation + apply preservesLimits_of_evaluation intro K change PreservesLimitsOfSize (yoneda.obj K) infer_instance -noncomputable instance yonedaFunctorReflectsLimits : +noncomputable instance yonedaFunctor_reflectsLimits : ReflectsLimitsOfSize.{t, w} (@yoneda C _) := inferInstance -noncomputable instance coyonedaFunctorReflectsLimits : +noncomputable instance coyonedaFunctor_reflectsLimits : ReflectsLimitsOfSize.{t, w} (@coyoneda C _) := inferInstance namespace Functor @@ -195,15 +195,15 @@ section Representable variable (F : Cᵒᵖ ⥤ Type v) [F.IsRepresentable] {J : Type*} [Category J] -noncomputable instance representablePreservesLimit (G : J ⥤ Cᵒᵖ) : +instance representable_preservesLimit (G : J ⥤ Cᵒᵖ) : PreservesLimit G F := - preservesLimitOfNatIso _ F.reprW + preservesLimit_of_natIso _ F.reprW variable (J) in -noncomputable instance representablePreservesLimitsOfShape : +instance representable_preservesLimitsOfShape : PreservesLimitsOfShape J F where -noncomputable instance representablePreservesLimits : +instance representable_preservesLimits : PreservesLimitsOfSize.{t, w} F where end Representable @@ -212,15 +212,15 @@ section Corepresentable variable (F : C ⥤ Type v) [F.IsCorepresentable] {J : Type*} [Category J] -noncomputable instance corepresentablePreservesLimit (G : J ⥤ C) : +instance corepresentable_preservesLimit (G : J ⥤ C) : PreservesLimit G F := - preservesLimitOfNatIso _ F.coreprW + preservesLimit_of_natIso _ F.coreprW variable (J) in -noncomputable instance corepresentablePreservesLimitsOfShape : +instance corepresentable_preservesLimitsOfShape : PreservesLimitsOfShape J F where -noncomputable instance corepresentablePreservesLimits : +instance corepresentable_preservesLimits : PreservesLimitsOfSize.{t, w} F where end Corepresentable diff --git a/Mathlib/CategoryTheory/Localization/FiniteProducts.lean b/Mathlib/CategoryTheory/Localization/FiniteProducts.lean index 7fe7c9bfe0e8a..d3be845da56a6 100644 --- a/Mathlib/CategoryTheory/Localization/FiniteProducts.lean +++ b/Mathlib/CategoryTheory/Localization/FiniteProducts.lean @@ -99,10 +99,10 @@ lemma hasProductsOfShape (J : Type) [Finite J] [HasProductsOfShape J C] /-- When `C` has finite products indexed by `J`, `W : MorphismProperty C` contains identities and is stable by products indexed by `J`, then any localization functor for `W` preserves finite products indexed by `J`. -/ -noncomputable def preservesProductsOfShape (J : Type) [Finite J] +lemma preservesProductsOfShape (J : Type) [Finite J] [HasProductsOfShape J C] (hW : W.IsStableUnderProductsOfShape J) : PreservesLimitsOfShape (Discrete J) L where - preservesLimit {F} := preservesLimitOfPreservesLimitCone (limit.isLimit F) + preservesLimit {F} := preservesLimit_of_preserves_limit_cone (limit.isLimit F) (HasProductsOfShapeAux.isLimitMapCone L hW F) variable [HasFiniteProducts C] [W.IsStableUnderFiniteProducts] @@ -112,10 +112,11 @@ lemma hasFiniteProducts : HasFiniteProducts D := ⟨fun _ => hasProductsOfShape L W _ (W.isStableUnderProductsOfShape_of_isStableUnderFiniteProducts _)⟩ +include W in /-- When `C` has finite products and `W : MorphismProperty C` contains identities and is stable by finite products, then any localization functor for `W` preserves finite products. -/ -noncomputable def preservesFiniteProducts : +lemma preservesFiniteProducts : PreservesFiniteProducts L where preserves J _ := preservesProductsOfShape L W J (W.isStableUnderProductsOfShape_of_isStableUnderFiniteProducts _) diff --git a/Mathlib/CategoryTheory/Monad/Comonadicity.lean b/Mathlib/CategoryTheory/Monad/Comonadicity.lean index 2649f2cfb69b3..d72e8b6e22780 100644 --- a/Mathlib/CategoryTheory/Monad/Comonadicity.lean +++ b/Mathlib/CategoryTheory/Monad/Comonadicity.lean @@ -223,10 +223,8 @@ Beck's comonadicity theorem, the converse is given in `comonadicOfCreatesFSplitE def createsFSplitEqualizersOfComonadic [ComonadicLeftAdjoint F] ⦃A B⦄ (f g : A ⟶ B) [F.IsCosplitPair f g] : CreatesLimit (parallelPair f g) F := by apply (config := {allowSynthFailures := true}) comonadicCreatesLimitOfPreservesLimit - · apply @preservesLimitOfIsoDiagram _ _ _ _ _ _ _ _ _ (diagramIsoParallelPair.{v₁} _).symm ?_ - dsimp - infer_instance - · apply @preservesLimitOfIsoDiagram _ _ _ _ _ _ _ _ _ (diagramIsoParallelPair.{v₁} _).symm ?_ + all_goals + apply @preservesLimit_of_iso_diagram _ _ _ _ _ _ _ _ _ (diagramIsoParallelPair.{v₁} _).symm ?_ dsimp infer_instance @@ -335,7 +333,7 @@ def comonadicOfHasPreservesFSplitEqualizersOfReflectsIsomorphisms [F.ReflectsIso ComonadicLeftAdjoint F := by have : ReflectsLimitOfIsCosplitPair F := ⟨fun f g _ => by have := HasEqualizerOfIsCosplitPair.out F f g - apply reflectsLimitOfReflectsIsomorphisms⟩ + apply reflectsLimit_of_reflectsIsomorphisms⟩ apply comonadicOfHasPreservesReflectsFSplitEqualizers adj end BeckComonadicity @@ -387,7 +385,7 @@ def comonadicOfHasPreservesCoreflexiveEqualizersOfReflectsIsomorphisms : · rw [← G.map_id] simp apply @unitEqualizerOfCoreflectsEqualizer _ _ _ _ _ _ _ _ ?_ - apply reflectsLimitOfReflectsIsomorphisms + apply reflectsLimit_of_reflectsIsomorphisms exact (comparisonAdjunction adj).toEquivalence.symm.isEquivalence_inverse end CoreflexiveComonadicity diff --git a/Mathlib/CategoryTheory/Monad/Limits.lean b/Mathlib/CategoryTheory/Monad/Limits.lean index 11d531a830a9b..8e99e8eace74d 100644 --- a/Mathlib/CategoryTheory/Monad/Limits.lean +++ b/Mathlib/CategoryTheory/Monad/Limits.lean @@ -154,7 +154,7 @@ Define the map `λ : TL ⟶ L`, which will serve as the structure of the coalgeb we will show is the colimiting object. We use the cocone constructed by `c` and the fact that `T` preserves colimits to produce this morphism. -/ -abbrev lambda : ((T : C ⥤ C).mapCocone c).pt ⟶ c.pt := +noncomputable abbrev lambda : ((T : C ⥤ C).mapCocone c).pt ⟶ c.pt := (isColimitOfPreserves _ t).desc (newCocone c) /-- (Impl) The key property defining the map `λ : TL ⟶ L`. -/ @@ -169,7 +169,7 @@ show it satisfies the two algebra laws, which follow from the algebra laws for t our `commuting` lemma. -/ @[simps] -def coconePoint : Algebra T where +noncomputable def coconePoint : Algebra T where A := c.pt a := lambda c t unit := by @@ -189,7 +189,7 @@ def coconePoint : Algebra T where /-- (Impl) Construct the lifted cocone in `Algebra T` which will be colimiting. -/ @[simps] -def liftedCocone : Cocone D where +noncomputable def liftedCocone : Cocone D where pt := coconePoint c t ι := { app := fun j => @@ -203,7 +203,7 @@ def liftedCocone : Cocone D where /-- (Impl) Prove that the lifted cocone is colimiting. -/ @[simps] -def liftedCoconeIsColimit : IsColimit (liftedCocone c t) where +noncomputable def liftedCoconeIsColimit : IsColimit (liftedCocone c t) where desc s := { f := t.desc ((forget T).mapCocone s) h := @@ -293,13 +293,13 @@ noncomputable def monadicCreatesColimitOfPreservesColimit (R : D ⥤ C) (K : J (Adjunction.toMonad (monadicAdjunction R))) (Adjunction.toMonad (monadicAdjunction R)).toFunctor := by dsimp - exact preservesColimitOfIsoDiagram _ i.symm + exact preservesColimit_of_iso_diagram _ i.symm letI : PreservesColimit (((K ⋙ A) ⋙ Monad.forget (Adjunction.toMonad (monadicAdjunction R))) ⋙ (Adjunction.toMonad (monadicAdjunction R)).toFunctor) (Adjunction.toMonad (monadicAdjunction R)).toFunctor := by dsimp - exact preservesColimitOfIsoDiagram _ (isoWhiskerRight i (monadicLeftAdjoint R ⋙ R)).symm + exact preservesColimit_of_iso_diagram _ (isoWhiskerRight i (monadicLeftAdjoint R ⋙ R)).symm letI : CreatesColimit (K ⋙ A) B := CategoryTheory.Monad.forgetCreatesColimit _ letI : CreatesColimit K (A ⋙ B) := CategoryTheory.compCreatesColimit _ _ let e := Monad.comparisonForget (monadicAdjunction R) @@ -309,9 +309,9 @@ noncomputable def monadicCreatesColimitOfPreservesColimit (R : D ⥤ C) (K : J noncomputable def monadicCreatesColimitsOfShapeOfPreservesColimitsOfShape (R : D ⥤ C) [MonadicRightAdjoint R] [PreservesColimitsOfShape J R] : CreatesColimitsOfShape J R := letI : PreservesColimitsOfShape J (monadicLeftAdjoint R) := by - apply (Adjunction.leftAdjointPreservesColimits (monadicAdjunction R)).1 + apply (Adjunction.leftAdjoint_preservesColimits (monadicAdjunction R)).1 letI : PreservesColimitsOfShape J (monadicLeftAdjoint R ⋙ R) := by - apply CategoryTheory.Limits.compPreservesColimitsOfShape _ _ + apply CategoryTheory.Limits.comp_preservesColimitsOfShape _ _ ⟨monadicCreatesColimitOfPreservesColimit _ _⟩ /-- A monadic functor creates colimits if it preserves colimits. -/ @@ -343,7 +343,7 @@ theorem hasColimitsOfShape_of_reflective (R : D ⥤ C) [Reflective R] [HasColimi has_colimit := fun F => by let c := (monadicLeftAdjoint R).mapCocone (colimit.cocone (F ⋙ R)) letI : PreservesColimitsOfShape J _ := - (monadicAdjunction R).leftAdjointPreservesColimits.1 + (monadicAdjunction R).leftAdjoint_preservesColimits.1 let t : IsColimit c := isColimitOfPreserves (monadicLeftAdjoint R) (colimit.isColimit _) apply HasColimit.mk ⟨_, (IsColimit.precomposeInvEquiv _ _).symm t⟩ apply @@ -357,7 +357,7 @@ theorem hasColimits_of_reflective (R : D ⥤ C) [Reflective R] [HasColimitsOfSiz /-- The reflector always preserves terminal objects. Note this in general doesn't apply to any other limit. -/ -noncomputable def leftAdjointPreservesTerminalOfReflective (R : D ⥤ C) [Reflective R] : +lemma leftAdjoint_preservesTerminal_of_reflective (R : D ⥤ C) [Reflective R] : PreservesLimitsOfShape (Discrete.{v} PEmpty) (monadicLeftAdjoint R) where preservesLimit {K} := by let F := Functor.empty.{v} D @@ -366,13 +366,14 @@ noncomputable def leftAdjointPreservesTerminalOfReflective (R : D ⥤ C) [Reflec intro c h haveI : HasLimit (F ⋙ R) := ⟨⟨⟨c, h⟩⟩⟩ haveI : HasLimit F := hasLimit_of_reflective F R + constructor apply isLimitChangeEmptyCone D (limit.isLimit F) apply (asIso ((monadicAdjunction R).counit.app _)).symm.trans apply (monadicLeftAdjoint R).mapIso letI := monadicCreatesLimits.{v, v} R - let A := CategoryTheory.preservesLimitOfCreatesLimitAndHasLimit F R - apply (A.preserves (limit.isLimit F)).conePointUniqueUpToIso h - apply preservesLimitOfIsoDiagram _ (Functor.emptyExt (F ⋙ R) _) + let A := CategoryTheory.preservesLimit_of_createsLimit_and_hasLimit F R + apply (isLimitOfPreserves _ (limit.isLimit F)).conePointUniqueUpToIso h + apply preservesLimit_of_iso_diagram _ (Functor.emptyExt (F ⋙ R) _) end @@ -484,7 +485,7 @@ Define the map `λ : L ⟶ TL`, which will serve as the structure of the algebra we will show is the limiting object. We use the cone constructed by `c` and the fact that `T` preserves limits to produce this morphism. -/ -abbrev lambda : c.pt ⟶ ((T : C ⥤ C).mapCone c).pt := +noncomputable abbrev lambda : c.pt ⟶ ((T : C ⥤ C).mapCone c).pt := (isLimitOfPreserves _ t).lift (newCone c) /-- (Impl) The key property defining the map `λ : L ⟶ TL`. -/ @@ -500,7 +501,7 @@ show it satisfies the two coalgebra laws, which follow from the coalgebra laws f and our `commuting` lemma. -/ @[simps] -def conePoint : Coalgebra T where +noncomputable def conePoint : Coalgebra T where A := c.pt a := lambda c t counit := t.hom_ext fun j ↦ by @@ -516,7 +517,7 @@ def conePoint : Coalgebra T where /-- (Impl) Construct the lifted cone in `Coalgebra T` which will be limiting. -/ @[simps] -def liftedCone : Cone D where +noncomputable def liftedCone : Cone D where pt := conePoint c t π := { app := fun j => @@ -530,7 +531,7 @@ def liftedCone : Cone D where /-- (Impl) Prove that the lifted cone is limiting. -/ @[simps] -def liftedConeIsLimit : IsLimit (liftedCone c t) where +noncomputable def liftedConeIsLimit : IsLimit (liftedCone c t) where lift s := { f := t.lift ((forget T).mapCone s) h := @@ -616,13 +617,13 @@ noncomputable def comonadicCreatesLimitOfPreservesLimit (R : D ⥤ C) (K : J ⥤ (Adjunction.toComonad (comonadicAdjunction R))) (Adjunction.toComonad (comonadicAdjunction R)).toFunctor := by dsimp - exact preservesLimitOfIsoDiagram _ i.symm + exact preservesLimit_of_iso_diagram _ i.symm letI : PreservesLimit (((K ⋙ A) ⋙ Comonad.forget (Adjunction.toComonad (comonadicAdjunction R))) ⋙ (Adjunction.toComonad (comonadicAdjunction R)).toFunctor) (Adjunction.toComonad (comonadicAdjunction R)).toFunctor := by dsimp - exact preservesLimitOfIsoDiagram _ (isoWhiskerRight i (comonadicRightAdjoint R ⋙ R)).symm + exact preservesLimit_of_iso_diagram _ (isoWhiskerRight i (comonadicRightAdjoint R ⋙ R)).symm letI : CreatesLimit (K ⋙ A) B := CategoryTheory.Comonad.forgetCreatesLimit _ letI : CreatesLimit K (A ⋙ B) := CategoryTheory.compCreatesLimit _ _ let e := Comonad.comparisonForget (comonadicAdjunction R) @@ -632,9 +633,9 @@ noncomputable def comonadicCreatesLimitOfPreservesLimit (R : D ⥤ C) (K : J ⥤ noncomputable def comonadicCreatesLimitsOfShapeOfPreservesLimitsOfShape (R : D ⥤ C) [ComonadicLeftAdjoint R] [PreservesLimitsOfShape J R] : CreatesLimitsOfShape J R := letI : PreservesLimitsOfShape J (comonadicRightAdjoint R) := by - apply (Adjunction.rightAdjointPreservesLimits (comonadicAdjunction R)).1 + apply (Adjunction.rightAdjoint_preservesLimits (comonadicAdjunction R)).1 letI : PreservesLimitsOfShape J (comonadicRightAdjoint R ⋙ R) := by - apply CategoryTheory.Limits.compPreservesLimitsOfShape _ _ + apply CategoryTheory.Limits.comp_preservesLimitsOfShape _ _ ⟨comonadicCreatesLimitOfPreservesLimit _ _⟩ /-- A comonadic functor creates limits if it preserves limits. -/ @@ -666,7 +667,7 @@ theorem hasLimitsOfShape_of_coreflective (R : D ⥤ C) [Coreflective R] [HasLimi has_limit := fun F => by let c := (comonadicRightAdjoint R).mapCone (limit.cone (F ⋙ R)) letI : PreservesLimitsOfShape J _ := - (comonadicAdjunction R).rightAdjointPreservesLimits.1 + (comonadicAdjunction R).rightAdjoint_preservesLimits.1 let t : IsLimit c := isLimitOfPreserves (comonadicRightAdjoint R) (limit.isLimit _) apply HasLimit.mk ⟨_, (IsLimit.postcomposeHomEquiv _ _).symm t⟩ apply @@ -680,7 +681,7 @@ theorem hasLimits_of_coreflective (R : D ⥤ C) [Coreflective R] [HasLimitsOfSiz /-- The coreflector always preserves initial objects. Note this in general doesn't apply to any other colimit. -/ -noncomputable def rightAdjointPreservesInitialOfCoreflective (R : D ⥤ C) [Coreflective R] : +lemma rightAdjoint_preservesInitial_of_coreflective (R : D ⥤ C) [Coreflective R] : PreservesColimitsOfShape (Discrete.{v} PEmpty) (comonadicRightAdjoint R) where preservesColimit {K} := by let F := Functor.empty.{v} D @@ -689,13 +690,14 @@ noncomputable def rightAdjointPreservesInitialOfCoreflective (R : D ⥤ C) [Core intro c h haveI : HasColimit (F ⋙ R) := ⟨⟨⟨c, h⟩⟩⟩ haveI : HasColimit F := hasColimit_of_coreflective F R + constructor apply isColimitChangeEmptyCocone D (colimit.isColimit F) apply (asIso ((comonadicAdjunction R).unit.app _)).trans apply (comonadicRightAdjoint R).mapIso letI := comonadicCreatesColimits.{v, v} R - let A := CategoryTheory.preservesColimitOfCreatesColimitAndHasColimit F R - apply (A.preserves (colimit.isColimit F)).coconePointUniqueUpToIso h - apply preservesColimitOfIsoDiagram _ (Functor.emptyExt (F ⋙ R) _) + let A := CategoryTheory.preservesColimit_of_createsColimit_and_hasColimit F R + apply (isColimitOfPreserves _ (colimit.isColimit F)).coconePointUniqueUpToIso h + apply preservesColimit_of_iso_diagram _ (Functor.emptyExt (F ⋙ R) _) end diff --git a/Mathlib/CategoryTheory/Monad/Monadicity.lean b/Mathlib/CategoryTheory/Monad/Monadicity.lean index dcb6b8d573feb..bd9b396629363 100644 --- a/Mathlib/CategoryTheory/Monad/Monadicity.lean +++ b/Mathlib/CategoryTheory/Monad/Monadicity.lean @@ -244,10 +244,8 @@ def createsGSplitCoequalizersOfMonadic [MonadicRightAdjoint G] ⦃A B⦄ (f g : [G.IsSplitPair f g] : CreatesColimit (parallelPair f g) G := by apply (config := {allowSynthFailures := true}) monadicCreatesColimitOfPreservesColimit -- Porting note: oddly (config := {allowSynthFailures := true}) had no effect here and below - · apply @preservesColimitOfIsoDiagram _ _ _ _ _ _ _ _ _ (diagramIsoParallelPair.{v₁} _).symm ?_ - dsimp - infer_instance - · apply @preservesColimitOfIsoDiagram _ _ _ _ _ _ _ _ _ (diagramIsoParallelPair.{v₁} _).symm ?_ + all_goals + apply @preservesColimit_of_iso_diagram _ _ _ _ _ _ _ _ _ (diagramIsoParallelPair.{v₁} _).symm ?_ dsimp infer_instance @@ -363,7 +361,7 @@ def monadicOfHasPreservesGSplitCoequalizersOfReflectsIsomorphisms [G.ReflectsIso MonadicRightAdjoint G := by have : ReflectsColimitOfIsSplitPair G := ⟨fun f g _ => by have := HasCoequalizerOfIsSplitPair.out G f g - apply reflectsColimitOfReflectsIsomorphisms⟩ + apply reflectsColimit_of_reflectsIsomorphisms⟩ apply monadicOfHasPreservesReflectsGSplitCoequalizers adj end BeckMonadicity @@ -409,7 +407,7 @@ def monadicOfHasPreservesReflexiveCoequalizersOfReflectsIsomorphisms : MonadicRi infer_instance -- Porting note: passing instances through apply @counitCoequalizerOfReflectsCoequalizer _ _ _ _ _ _ _ _ ?_ - apply reflectsColimitOfReflectsIsomorphisms + apply reflectsColimit_of_reflectsIsomorphisms exact (comparisonAdjunction adj).toEquivalence.isEquivalence_inverse end ReflexiveMonadicity diff --git a/Mathlib/CategoryTheory/Monoidal/Internal/Limits.lean b/Mathlib/CategoryTheory/Monoidal/Internal/Limits.lean index 00f12f9aee66e..84cdce3c91ae3 100644 --- a/Mathlib/CategoryTheory/Monoidal/Internal/Limits.lean +++ b/Mathlib/CategoryTheory/Monoidal/Internal/Limits.lean @@ -79,9 +79,9 @@ instance hasLimitsOfShape [HasLimitsOfShape J C] : HasLimitsOfShape J (Mon_ C) w { cone := limitCone F isLimit := limitConeIsLimit F } -instance forgetPreservesLimitsOfShape : PreservesLimitsOfShape J (Mon_.forget C) where +instance forget_freservesLimitsOfShape : PreservesLimitsOfShape J (Mon_.forget C) where preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone (limitConeIsLimit F) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit F) (IsLimit.ofIsoLimit (limit.isLimit (F ⋙ Mon_.forget C)) (forgetMapConeLimitConeIso F).symm) end Mon_ diff --git a/Mathlib/CategoryTheory/Monoidal/Preadditive.lean b/Mathlib/CategoryTheory/Monoidal/Preadditive.lean index 5be7683ef2df1..d7d052e314259 100644 --- a/Mathlib/CategoryTheory/Monoidal/Preadditive.lean +++ b/Mathlib/CategoryTheory/Monoidal/Preadditive.lean @@ -116,20 +116,20 @@ theorem sum_tensor {P Q R S : C} {J : Type*} (s : Finset J) (f : P ⟶ Q) (g : J instance (X : C) : PreservesFiniteBiproducts (tensorLeft X) where preserves {J} := { preserves := fun {f} => - { preserves := fun {b} i => isBilimitOfTotal _ (by + { preserves := fun {b} i => ⟨isBilimitOfTotal _ (by dsimp simp_rw [← id_tensorHom] simp only [← tensor_comp, Category.comp_id, ← tensor_sum, ← tensor_id, - IsBilimit.total i]) } } + IsBilimit.total i])⟩ } } instance (X : C) : PreservesFiniteBiproducts (tensorRight X) where preserves {J} := { preserves := fun {f} => - { preserves := fun {b} i => isBilimitOfTotal _ (by + { preserves := fun {b} i => ⟨isBilimitOfTotal _ (by dsimp simp_rw [← tensorHom_id] simp only [← tensor_comp, Category.comp_id, ← sum_tensor, ← tensor_id, - IsBilimit.total i]) } } + IsBilimit.total i])⟩ } } variable [HasFiniteBiproducts C] diff --git a/Mathlib/CategoryTheory/Preadditive/AdditiveFunctor.lean b/Mathlib/CategoryTheory/Preadditive/AdditiveFunctor.lean index a5d134fdcfce6..4a0fd32e2e836 100644 --- a/Mathlib/CategoryTheory/Preadditive/AdditiveFunctor.lean +++ b/Mathlib/CategoryTheory/Preadditive/AdditiveFunctor.lean @@ -157,9 +157,9 @@ instance (priority := 100) preservesFiniteBiproductsOfAdditive [Additive F] : preserves := { preserves := { preserves := fun hb => - isBilimitOfTotal _ (by + ⟨isBilimitOfTotal _ (by simp_rw [F.mapBicone_π, F.mapBicone_ι, ← F.map_comp] - erw [← F.map_sum, ← F.map_id, IsBilimit.total hb])} } + erw [← F.map_sum, ← F.map_id, IsBilimit.total hb])⟩ } } theorem additive_of_preservesBinaryBiproducts [HasBinaryBiproducts C] [PreservesZeroMorphisms F] [PreservesBinaryBiproducts F] : Additive F where @@ -172,7 +172,7 @@ lemma additive_of_preserves_binary_products [HasBinaryProducts C] [PreservesLimitsOfShape (Discrete WalkingPair) F] [F.PreservesZeroMorphisms] : F.Additive := by have : HasBinaryBiproducts C := HasBinaryBiproducts.of_hasBinaryProducts - have := preservesBinaryBiproductsOfPreservesBinaryProducts F + have := preservesBinaryBiproducts_of_preservesBinaryProducts F exact Functor.additive_of_preservesBinaryBiproducts F end @@ -254,9 +254,9 @@ variable [Preadditive D] [HasZeroObject C] [HasZeroObject D] [HasBinaryBiproduct section -attribute [local instance] preservesBinaryBiproductsOfPreservesBinaryProducts +attribute [local instance] preservesBinaryBiproducts_of_preservesBinaryProducts -attribute [local instance] preservesBinaryBiproductsOfPreservesBinaryCoproducts +attribute [local instance] preservesBinaryBiproducts_of_preservesBinaryCoproducts /-- Turn a left exact functor into an additive functor. -/ def AdditiveFunctor.ofLeftExact : (C ⥤ₗ D) ⥤ C ⥤+ D := diff --git a/Mathlib/CategoryTheory/Preadditive/Biproducts.lean b/Mathlib/CategoryTheory/Preadditive/Biproducts.lean index 5b06f3702ea3b..5cbeb420311dc 100644 --- a/Mathlib/CategoryTheory/Preadditive/Biproducts.lean +++ b/Mathlib/CategoryTheory/Preadditive/Biproducts.lean @@ -825,40 +825,40 @@ variable {J : Type} [Fintype J] /-- A functor between preadditive categories that preserves (zero morphisms and) finite biproducts preserves finite products. -/ -def preservesProductOfPreservesBiproduct {f : J → C} [PreservesBiproduct f F] : +lemma preservesProduct_of_preservesBiproduct {f : J → C} [PreservesBiproduct f F] : PreservesLimit (Discrete.functor f) F where preserves hc := - IsLimit.ofIsoLimit + ⟨IsLimit.ofIsoLimit ((IsLimit.postcomposeInvEquiv (Discrete.compNatIsoDiscrete _ _) _).symm (isBilimitOfPreserves F (biconeIsBilimitOfLimitConeOfIsLimit hc)).isLimit) <| - Cones.ext (Iso.refl _) (by rintro ⟨⟩; simp) + Cones.ext (Iso.refl _) (by rintro ⟨⟩; simp)⟩ section -attribute [local instance] preservesProductOfPreservesBiproduct +attribute [local instance] preservesProduct_of_preservesBiproduct /-- A functor between preadditive categories that preserves (zero morphisms and) finite biproducts preserves finite products. -/ -def preservesProductsOfShapeOfPreservesBiproductsOfShape [PreservesBiproductsOfShape J F] : +lemma preservesProductsOfShape_of_preservesBiproductsOfShape [PreservesBiproductsOfShape J F] : PreservesLimitsOfShape (Discrete J) F where - preservesLimit {_} := preservesLimitOfIsoDiagram _ Discrete.natIsoFunctor.symm + preservesLimit {_} := preservesLimit_of_iso_diagram _ Discrete.natIsoFunctor.symm end /-- A functor between preadditive categories that preserves (zero morphisms and) finite products preserves finite biproducts. -/ -def preservesBiproductOfPreservesProduct {f : J → C} [PreservesLimit (Discrete.functor f) F] : +lemma preservesBiproduct_of_preservesProduct {f : J → C} [PreservesLimit (Discrete.functor f) F] : PreservesBiproduct f F where preserves {b} hb := - isBilimitOfIsLimit _ <| + ⟨isBilimitOfIsLimit _ <| IsLimit.ofIsoLimit ((IsLimit.postcomposeHomEquiv (Discrete.compNatIsoDiscrete _ _) (F.mapCone b.toCone)).symm (isLimitOfPreserves F hb.isLimit)) <| - Cones.ext (Iso.refl _) (by rintro ⟨⟩; simp) + Cones.ext (Iso.refl _) (by rintro ⟨⟩; simp)⟩ /-- If the (product-like) biproduct comparison for `F` and `f` is a monomorphism, then `F` preserves the biproduct of `f`. For the converse, see `mapBiproduct`. -/ -def preservesBiproductOfMonoBiproductComparison {f : J → C} [HasBiproduct f] +lemma preservesBiproduct_of_mono_biproductComparison {f : J → C} [HasBiproduct f] [HasBiproduct (F.obj ∘ f)] [Mono (biproductComparison F f)] : PreservesBiproduct f F := by haveI : HasProduct fun b => F.obj (f b) := by change HasProduct (F.obj ∘ f) @@ -872,101 +872,104 @@ def preservesBiproductOfMonoBiproductComparison {f : J → C} [HasBiproduct f] haveI : IsIso (piComparison F f) := by rw [that] infer_instance - haveI := PreservesProduct.ofIsoComparison F f - apply preservesBiproductOfPreservesProduct + haveI := PreservesProduct.of_iso_comparison F f + apply preservesBiproduct_of_preservesProduct /-- If the (coproduct-like) biproduct comparison for `F` and `f` is an epimorphism, then `F` preserves the biproduct of `F` and `f`. For the converse, see `mapBiproduct`. -/ -def preservesBiproductOfEpiBiproductComparison' {f : J → C} [HasBiproduct f] +lemma preservesBiproduct_of_epi_biproductComparison' {f : J → C} [HasBiproduct f] [HasBiproduct (F.obj ∘ f)] [Epi (biproductComparison' F f)] : PreservesBiproduct f F := by haveI : Epi (splitEpiBiproductComparison F f).section_ := by simpa haveI : IsIso (biproductComparison F f) := IsIso.of_epi_section' (splitEpiBiproductComparison F f) - apply preservesBiproductOfMonoBiproductComparison + apply preservesBiproduct_of_mono_biproductComparison /-- A functor between preadditive categories that preserves (zero morphisms and) finite products preserves finite biproducts. -/ -def preservesBiproductsOfShapeOfPreservesProductsOfShape [PreservesLimitsOfShape (Discrete J) F] : +lemma preservesBiproductsOfShape_of_preservesProductsOfShape + [PreservesLimitsOfShape (Discrete J) F] : PreservesBiproductsOfShape J F where - preserves {_} := preservesBiproductOfPreservesProduct F + preserves {_} := preservesBiproduct_of_preservesProduct F /-- A functor between preadditive categories that preserves (zero morphisms and) finite biproducts preserves finite coproducts. -/ -def preservesCoproductOfPreservesBiproduct {f : J → C} [PreservesBiproduct f F] : +lemma preservesCoproduct_of_preservesBiproduct {f : J → C} [PreservesBiproduct f F] : PreservesColimit (Discrete.functor f) F where preserves {c} hc := - IsColimit.ofIsoColimit + ⟨IsColimit.ofIsoColimit ((IsColimit.precomposeHomEquiv (Discrete.compNatIsoDiscrete _ _) _).symm (isBilimitOfPreserves F (biconeIsBilimitOfColimitCoconeOfIsColimit hc)).isColimit) <| - Cocones.ext (Iso.refl _) (by rintro ⟨⟩; simp) + Cocones.ext (Iso.refl _) (by rintro ⟨⟩; simp)⟩ section -attribute [local instance] preservesCoproductOfPreservesBiproduct +attribute [local instance] preservesCoproduct_of_preservesBiproduct /-- A functor between preadditive categories that preserves (zero morphisms and) finite biproducts preserves finite coproducts. -/ -def preservesCoproductsOfShapeOfPreservesBiproductsOfShape [PreservesBiproductsOfShape J F] : +lemma preservesCoproductsOfShape_of_preservesBiproductsOfShape [PreservesBiproductsOfShape J F] : PreservesColimitsOfShape (Discrete J) F where - preservesColimit {_} := preservesColimitOfIsoDiagram _ Discrete.natIsoFunctor.symm + preservesColimit {_} := preservesColimit_of_iso_diagram _ Discrete.natIsoFunctor.symm end /-- A functor between preadditive categories that preserves (zero morphisms and) finite coproducts preserves finite biproducts. -/ -def preservesBiproductOfPreservesCoproduct {f : J → C} [PreservesColimit (Discrete.functor f) F] : +lemma preservesBiproduct_of_preservesCoproduct {f : J → C} + [PreservesColimit (Discrete.functor f) F] : PreservesBiproduct f F where preserves {b} hb := - isBilimitOfIsColimit _ <| + ⟨isBilimitOfIsColimit _ <| IsColimit.ofIsoColimit ((IsColimit.precomposeInvEquiv (Discrete.compNatIsoDiscrete _ _) (F.mapCocone b.toCocone)).symm (isColimitOfPreserves F hb.isColimit)) <| - Cocones.ext (Iso.refl _) (by rintro ⟨⟩; simp) + Cocones.ext (Iso.refl _) (by rintro ⟨⟩; simp)⟩ /-- A functor between preadditive categories that preserves (zero morphisms and) finite coproducts preserves finite biproducts. -/ -def preservesBiproductsOfShapeOfPreservesCoproductsOfShape +lemma preservesBiproductsOfShape_of_preservesCoproductsOfShape [PreservesColimitsOfShape (Discrete J) F] : PreservesBiproductsOfShape J F where - preserves {_} := preservesBiproductOfPreservesCoproduct F + preserves {_} := preservesBiproduct_of_preservesCoproduct F end Fintype /-- A functor between preadditive categories that preserves (zero morphisms and) binary biproducts preserves binary products. -/ -def preservesBinaryProductOfPreservesBinaryBiproduct {X Y : C} [PreservesBinaryBiproduct X Y F] : +lemma preservesBinaryProduct_of_preservesBinaryBiproduct {X Y : C} + [PreservesBinaryBiproduct X Y F] : PreservesLimit (pair X Y) F where - preserves {c} hc := IsLimit.ofIsoLimit + preserves {c} hc := ⟨IsLimit.ofIsoLimit ((IsLimit.postcomposeInvEquiv (diagramIsoPair _) _).symm (isBinaryBilimitOfPreserves F (binaryBiconeIsBilimitOfLimitConeOfIsLimit hc)).isLimit) <| Cones.ext (by dsimp; rfl) fun j => by - rcases j with ⟨⟨⟩⟩ <;> simp + rcases j with ⟨⟨⟩⟩ <;> simp⟩ section -attribute [local instance] preservesBinaryProductOfPreservesBinaryBiproduct +attribute [local instance] preservesBinaryProduct_of_preservesBinaryBiproduct /-- A functor between preadditive categories that preserves (zero morphisms and) binary biproducts preserves binary products. -/ -def preservesBinaryProductsOfPreservesBinaryBiproducts [PreservesBinaryBiproducts F] : +lemma preservesBinaryProducts_of_preservesBinaryBiproducts [PreservesBinaryBiproducts F] : PreservesLimitsOfShape (Discrete WalkingPair) F where - preservesLimit {_} := preservesLimitOfIsoDiagram _ (diagramIsoPair _).symm + preservesLimit {_} := preservesLimit_of_iso_diagram _ (diagramIsoPair _).symm end /-- A functor between preadditive categories that preserves (zero morphisms and) binary products preserves binary biproducts. -/ -def preservesBinaryBiproductOfPreservesBinaryProduct {X Y : C} [PreservesLimit (pair X Y) F] : +lemma preservesBinaryBiproduct_of_preservesBinaryProduct {X Y : C} [PreservesLimit (pair X Y) F] : PreservesBinaryBiproduct X Y F where - preserves {b} hb := isBinaryBilimitOfIsLimit _ <| IsLimit.ofIsoLimit + preserves {b} hb := ⟨isBinaryBilimitOfIsLimit _ <| IsLimit.ofIsoLimit ((IsLimit.postcomposeHomEquiv (diagramIsoPair _) (F.mapCone b.toCone)).symm (isLimitOfPreserves F hb.isLimit)) <| Cones.ext (by dsimp; rfl) fun j => by - rcases j with ⟨⟨⟩⟩ <;> simp + rcases j with ⟨⟨⟩⟩ <;> simp⟩ /-- If the (product-like) biproduct comparison for `F`, `X` and `Y` is a monomorphism, then `F` preserves the biproduct of `X` and `Y`. For the converse, see `map_biprod`. -/ -def preservesBinaryBiproductOfMonoBiprodComparison {X Y : C} [HasBinaryBiproduct X Y] +lemma preservesBinaryBiproduct_of_mono_biprodComparison {X Y : C} [HasBinaryBiproduct X Y] [HasBinaryBiproduct (F.obj X) (F.obj Y)] [Mono (biprodComparison F X Y)] : PreservesBinaryBiproduct X Y F := by have that : @@ -977,66 +980,68 @@ def preservesBinaryBiproductOfMonoBiprodComparison {X Y : C} [HasBinaryBiproduct haveI : IsIso (prodComparison F X Y) := by rw [that] infer_instance - haveI := PreservesLimitPair.ofIsoProdComparison F X Y - apply preservesBinaryBiproductOfPreservesBinaryProduct + haveI := PreservesLimitPair.of_iso_prod_comparison F X Y + apply preservesBinaryBiproduct_of_preservesBinaryProduct /-- If the (coproduct-like) biproduct comparison for `F`, `X` and `Y` is an epimorphism, then `F` preserves the biproduct of `X` and `Y`. For the converse, see `mapBiprod`. -/ -def preservesBinaryBiproductOfEpiBiprodComparison' {X Y : C} [HasBinaryBiproduct X Y] +lemma preservesBinaryBiproduct_of_epi_biprodComparison' {X Y : C} [HasBinaryBiproduct X Y] [HasBinaryBiproduct (F.obj X) (F.obj Y)] [Epi (biprodComparison' F X Y)] : PreservesBinaryBiproduct X Y F := by haveI : Epi (splitEpiBiprodComparison F X Y).section_ := by simpa haveI : IsIso (biprodComparison F X Y) := IsIso.of_epi_section' (splitEpiBiprodComparison F X Y) - apply preservesBinaryBiproductOfMonoBiprodComparison + apply preservesBinaryBiproduct_of_mono_biprodComparison /-- A functor between preadditive categories that preserves (zero morphisms and) binary products preserves binary biproducts. -/ -def preservesBinaryBiproductsOfPreservesBinaryProducts +lemma preservesBinaryBiproducts_of_preservesBinaryProducts [PreservesLimitsOfShape (Discrete WalkingPair) F] : PreservesBinaryBiproducts F where - preserves {_} {_} := preservesBinaryBiproductOfPreservesBinaryProduct F + preserves {_} {_} := preservesBinaryBiproduct_of_preservesBinaryProduct F /-- A functor between preadditive categories that preserves (zero morphisms and) binary biproducts preserves binary coproducts. -/ -def preservesBinaryCoproductOfPreservesBinaryBiproduct {X Y : C} [PreservesBinaryBiproduct X Y F] : +lemma preservesBinaryCoproduct_of_preservesBinaryBiproduct {X Y : C} + [PreservesBinaryBiproduct X Y F] : PreservesColimit (pair X Y) F where preserves {c} hc := - IsColimit.ofIsoColimit + ⟨IsColimit.ofIsoColimit ((IsColimit.precomposeHomEquiv (diagramIsoPair _) _).symm (isBinaryBilimitOfPreserves F (binaryBiconeIsBilimitOfColimitCoconeOfIsColimit hc)).isColimit) <| Cocones.ext (by dsimp; rfl) fun j => by - rcases j with ⟨⟨⟩⟩ <;> simp + rcases j with ⟨⟨⟩⟩ <;> simp⟩ section -attribute [local instance] preservesBinaryCoproductOfPreservesBinaryBiproduct +attribute [local instance] preservesBinaryCoproduct_of_preservesBinaryBiproduct /-- A functor between preadditive categories that preserves (zero morphisms and) binary biproducts preserves binary coproducts. -/ -def preservesBinaryCoproductsOfPreservesBinaryBiproducts [PreservesBinaryBiproducts F] : +lemma preservesBinaryCoproducts_of_preservesBinaryBiproducts [PreservesBinaryBiproducts F] : PreservesColimitsOfShape (Discrete WalkingPair) F where - preservesColimit {_} := preservesColimitOfIsoDiagram _ (diagramIsoPair _).symm + preservesColimit {_} := preservesColimit_of_iso_diagram _ (diagramIsoPair _).symm end /-- A functor between preadditive categories that preserves (zero morphisms and) binary coproducts preserves binary biproducts. -/ -def preservesBinaryBiproductOfPreservesBinaryCoproduct {X Y : C} [PreservesColimit (pair X Y) F] : +lemma preservesBinaryBiproduct_of_preservesBinaryCoproduct {X Y : C} + [PreservesColimit (pair X Y) F] : PreservesBinaryBiproduct X Y F where preserves {b} hb := - isBinaryBilimitOfIsColimit _ <| + ⟨isBinaryBilimitOfIsColimit _ <| IsColimit.ofIsoColimit ((IsColimit.precomposeInvEquiv (diagramIsoPair _) (F.mapCocone b.toCocone)).symm (isColimitOfPreserves F hb.isColimit)) <| Cocones.ext (Iso.refl _) fun j => by - rcases j with ⟨⟨⟩⟩ <;> simp + rcases j with ⟨⟨⟩⟩ <;> simp⟩ /-- A functor between preadditive categories that preserves (zero morphisms and) binary coproducts preserves binary biproducts. -/ -def preservesBinaryBiproductsOfPreservesBinaryCoproducts +lemma preservesBinaryBiproducts_of_preservesBinaryCoproducts [PreservesColimitsOfShape (Discrete WalkingPair) F] : PreservesBinaryBiproducts F where - preserves {_} {_} := preservesBinaryBiproductOfPreservesBinaryCoproduct F + preserves {_} {_} := preservesBinaryBiproduct_of_preservesBinaryCoproduct F end Limits diff --git a/Mathlib/CategoryTheory/Preadditive/Injective.lean b/Mathlib/CategoryTheory/Preadditive/Injective.lean index 4ee0d211f19f3..d9ce92c4eef09 100644 --- a/Mathlib/CategoryTheory/Preadditive/Injective.lean +++ b/Mathlib/CategoryTheory/Preadditive/Injective.lean @@ -270,7 +270,7 @@ theorem injective_of_map_injective (adj : F ⊣ G) [G.Full] [G.Faithful] (I : D) (hI : Injective (G.obj I)) : Injective I := ⟨fun {X} {Y} f g => by intro - haveI : PreservesLimitsOfSize.{0, 0} G := adj.rightAdjointPreservesLimits + haveI : PreservesLimitsOfSize.{0, 0} G := adj.rightAdjoint_preservesLimits rcases hI.factors (G.map f) (G.map g) with ⟨w,h⟩ use inv (adj.counit.app _) ≫ F.map w ≫ adj.counit.app _ exact G.map_injective (by simpa)⟩ @@ -283,7 +283,7 @@ def mapInjectivePresentation (adj : F ⊣ G) [F.PreservesMonomorphisms] (X : D) injective := adj.map_injective _ I.injective f := G.map I.f mono := by - haveI : PreservesLimitsOfSize.{0, 0} G := adj.rightAdjointPreservesLimits; infer_instance + haveI : PreservesLimitsOfSize.{0, 0} G := adj.rightAdjoint_preservesLimits; infer_instance /-- Given an adjunction `F ⊣ G` such that `F` preserves monomorphisms and is faithful, then any injective presentation of `F(X)` can be pulled back to an injective presentation of `X`. diff --git a/Mathlib/CategoryTheory/Preadditive/LeftExact.lean b/Mathlib/CategoryTheory/Preadditive/LeftExact.lean index 1df5546b1d7c7..750d8578b60ca 100644 --- a/Mathlib/CategoryTheory/Preadditive/LeftExact.lean +++ b/Mathlib/CategoryTheory/Preadditive/LeftExact.lean @@ -55,36 +55,38 @@ def isLimitMapConeBinaryFanOfPreservesKernels {X Y Z : C} (π₁ : Z ⟶ X) (π (isLimitMapConeForkEquiv' F bc.inl_snd (isLimitOfPreserves F hf))).isLimit /-- A kernel preserving functor between preadditive categories preserves any pair being a limit. -/ -def preservesBinaryProductOfPreservesKernels +lemma preservesBinaryProduct_of_preservesKernels [∀ {X Y} (f : X ⟶ Y), PreservesLimit (parallelPair f 0) F] {X Y : C} : PreservesLimit (pair X Y) F where preserves {c} hc := - IsLimit.ofIsoLimit + ⟨IsLimit.ofIsoLimit (isLimitMapConeBinaryFanOfPreservesKernels F _ _ (IsLimit.ofIsoLimit hc (isoBinaryFanMk c))) - ((Cones.functoriality _ F).mapIso (isoBinaryFanMk c).symm) + ((Cones.functoriality _ F).mapIso (isoBinaryFanMk c).symm)⟩ -attribute [local instance] preservesBinaryProductOfPreservesKernels +attribute [local instance] preservesBinaryProduct_of_preservesKernels /-- A kernel preserving functor between preadditive categories preserves binary products. -/ -def preservesBinaryProductsOfPreservesKernels +lemma preservesBinaryProducts_of_preservesKernels [∀ {X Y} (f : X ⟶ Y), PreservesLimit (parallelPair f 0) F] : PreservesLimitsOfShape (Discrete WalkingPair) F where - preservesLimit := preservesLimitOfIsoDiagram F (diagramIsoPair _).symm + preservesLimit := preservesLimit_of_iso_diagram F (diagramIsoPair _).symm -attribute [local instance] preservesBinaryProductsOfPreservesKernels +attribute [local instance] preservesBinaryProducts_of_preservesKernels variable [HasBinaryBiproducts C] /-- A functor between preadditive categories preserves the equalizer of two morphisms if it preserves all kernels. -/ -def preservesEqualizerOfPreservesKernels [∀ {X Y} (f : X ⟶ Y), PreservesLimit (parallelPair f 0) F] +lemma preservesEqualizer_of_preservesKernels + [∀ {X Y} (f : X ⟶ Y), PreservesLimit (parallelPair f 0) F] {X Y : C} (f g : X ⟶ Y) : PreservesLimit (parallelPair f g) F := by - letI := preservesBinaryBiproductsOfPreservesBinaryProducts F + letI := preservesBinaryBiproducts_of_preservesBinaryProducts F haveI := additive_of_preservesBinaryBiproducts F constructor; intro c i let c' := isLimitKernelForkOfFork (i.ofIsoLimit (Fork.isoForkOfι c)) dsimp only [kernelForkOfFork_ofι] at c' let iFc := isLimitForkMapOfIsLimit' F _ c' + constructor apply IsLimit.ofIsoLimit _ ((Cones.functoriality _ F).mapIso (Fork.isoForkOfι c).symm) apply (isLimitMapConeForkEquiv F (Fork.condition c)).invFun let p : parallelPair (F.map (f - g)) 0 ≅ parallelPair (F.map f - F.map g) 0 := @@ -96,24 +98,24 @@ def preservesEqualizerOfPreservesKernels [∀ {X Y} (f : X ⟶ Y), PreservesLimi /-- A functor between preadditive categories preserves all equalizers if it preserves all kernels. -/ -def preservesEqualizersOfPreservesKernels +lemma preservesEqualizers_of_preservesKernels [∀ {X Y} (f : X ⟶ Y), PreservesLimit (parallelPair f 0) F] : PreservesLimitsOfShape WalkingParallelPair F where preservesLimit {K} := by - letI := preservesEqualizerOfPreservesKernels F (K.map WalkingParallelPairHom.left) + letI := preservesEqualizer_of_preservesKernels F (K.map WalkingParallelPairHom.left) (K.map WalkingParallelPairHom.right) - apply preservesLimitOfIsoDiagram F (diagramIsoParallelPair K).symm + apply preservesLimit_of_iso_diagram F (diagramIsoParallelPair K).symm /-- A functor between preadditive categories which preserves kernels preserves all finite limits. -/ -def preservesFiniteLimitsOfPreservesKernels [HasFiniteProducts C] [HasEqualizers C] +lemma preservesFiniteLimits_of_preservesKernels [HasFiniteProducts C] [HasEqualizers C] [HasZeroObject C] [HasZeroObject D] [∀ {X Y} (f : X ⟶ Y), PreservesLimit (parallelPair f 0) F] : PreservesFiniteLimits F := by - letI := preservesEqualizersOfPreservesKernels F - letI := preservesTerminalObjectOfPreservesZeroMorphisms F - letI := preservesLimitsOfShapePemptyOfPreservesTerminal F - letI : PreservesFiniteProducts F := ⟨preservesFiniteProductsOfPreservesBinaryAndTerminal F⟩ - exact preservesFiniteLimitsOfPreservesEqualizersAndFiniteProducts F + letI := preservesEqualizers_of_preservesKernels F + letI := preservesTerminalObject_of_preservesZeroMorphisms F + letI := preservesLimitsOfShape_pempty_of_preservesTerminal F + letI : PreservesFiniteProducts F := ⟨preservesFiniteProducts_of_preserves_binary_and_terminal F⟩ + exact preservesFiniteLimits_of_preservesEqualizers_and_finiteProducts F end FiniteLimits @@ -134,39 +136,40 @@ def isColimitMapCoconeBinaryCofanOfPreservesCokernels {X Y Z : C} (ι₁ : X ⟶ /-- A cokernel preserving functor between preadditive categories preserves any pair being a colimit. -/ -def preservesCoproductOfPreservesCokernels +lemma preservesCoproduct_of_preservesCokernels [∀ {X Y} (f : X ⟶ Y), PreservesColimit (parallelPair f 0) F] {X Y : C} : PreservesColimit (pair X Y) F where preserves {c} hc := - IsColimit.ofIsoColimit + ⟨IsColimit.ofIsoColimit (isColimitMapCoconeBinaryCofanOfPreservesCokernels F _ _ (IsColimit.ofIsoColimit hc (isoBinaryCofanMk c))) - ((Cocones.functoriality _ F).mapIso (isoBinaryCofanMk c).symm) + ((Cocones.functoriality _ F).mapIso (isoBinaryCofanMk c).symm)⟩ -attribute [local instance] preservesCoproductOfPreservesCokernels +attribute [local instance] preservesCoproduct_of_preservesCokernels /-- A cokernel preserving functor between preadditive categories preserves binary coproducts. -/ -def preservesBinaryCoproductsOfPreservesCokernels +lemma preservesBinaryCoproducts_of_preservesCokernels [∀ {X Y} (f : X ⟶ Y), PreservesColimit (parallelPair f 0) F] : PreservesColimitsOfShape (Discrete WalkingPair) F where - preservesColimit := preservesColimitOfIsoDiagram F (diagramIsoPair _).symm + preservesColimit := preservesColimit_of_iso_diagram F (diagramIsoPair _).symm -attribute [local instance] preservesBinaryCoproductsOfPreservesCokernels +attribute [local instance] preservesBinaryCoproducts_of_preservesCokernels variable [HasBinaryBiproducts C] /-- A functor between preadditive categories preserves the coequalizer of two morphisms if it preserves all cokernels. -/ -def preservesCoequalizerOfPreservesCokernels +lemma preservesCoequalizer_of_preservesCokernels [∀ {X Y} (f : X ⟶ Y), PreservesColimit (parallelPair f 0) F] {X Y : C} (f g : X ⟶ Y) : PreservesColimit (parallelPair f g) F := by - letI := preservesBinaryBiproductsOfPreservesBinaryCoproducts F + letI := preservesBinaryBiproducts_of_preservesBinaryCoproducts F haveI := additive_of_preservesBinaryBiproducts F constructor intro c i let c' := isColimitCokernelCoforkOfCofork (i.ofIsoColimit (Cofork.isoCoforkOfπ c)) dsimp only [cokernelCoforkOfCofork_ofπ] at c' let iFc := isColimitCoforkMapOfIsColimit' F _ c' + constructor apply IsColimit.ofIsoColimit _ ((Cocones.functoriality _ F).mapIso (Cofork.isoCoforkOfπ c).symm) apply (isColimitMapCoconeCoforkEquiv F (Cofork.condition c)).invFun @@ -179,24 +182,24 @@ def preservesCoequalizerOfPreservesCokernels /-- A functor between preadditive categories preserves all coequalizers if it preserves all kernels. -/ -def preservesCoequalizersOfPreservesCokernels +lemma preservesCoequalizers_of_preservesCokernels [∀ {X Y} (f : X ⟶ Y), PreservesColimit (parallelPair f 0) F] : PreservesColimitsOfShape WalkingParallelPair F where preservesColimit {K} := by - letI := preservesCoequalizerOfPreservesCokernels F (K.map Limits.WalkingParallelPairHom.left) + letI := preservesCoequalizer_of_preservesCokernels F (K.map Limits.WalkingParallelPairHom.left) (K.map Limits.WalkingParallelPairHom.right) - apply preservesColimitOfIsoDiagram F (diagramIsoParallelPair K).symm + apply preservesColimit_of_iso_diagram F (diagramIsoParallelPair K).symm /-- A functor between preadditive categories which preserves kernels preserves all finite limits. -/ -def preservesFiniteColimitsOfPreservesCokernels [HasFiniteCoproducts C] [HasCoequalizers C] +lemma preservesFiniteColimits_of_preservesCokernels [HasFiniteCoproducts C] [HasCoequalizers C] [HasZeroObject C] [HasZeroObject D] [∀ {X Y} (f : X ⟶ Y), PreservesColimit (parallelPair f 0) F] : PreservesFiniteColimits F := by - letI := preservesCoequalizersOfPreservesCokernels F - letI := preservesInitialObjectOfPreservesZeroMorphisms F - letI := preservesColimitsOfShapePemptyOfPreservesInitial F + letI := preservesCoequalizers_of_preservesCokernels F + letI := preservesInitialObject_of_preservesZeroMorphisms F + letI := preservesColimitsOfShape_pempty_of_preservesInitial F letI : PreservesFiniteCoproducts F := ⟨preservesFiniteCoproductsOfPreservesBinaryAndInitial F⟩ - exact preservesFiniteColimitsOfPreservesCoequalizersAndFiniteCoproducts F + exact preservesFiniteColimits_of_preservesCoequalizers_and_finiteCoproducts F end FiniteColimits diff --git a/Mathlib/CategoryTheory/Preadditive/Projective.lean b/Mathlib/CategoryTheory/Preadditive/Projective.lean index 813b5d697a416..b576bee9c0e8a 100644 --- a/Mathlib/CategoryTheory/Preadditive/Projective.lean +++ b/Mathlib/CategoryTheory/Preadditive/Projective.lean @@ -199,7 +199,7 @@ theorem map_projective (adj : F ⊣ G) [G.PreservesEpimorphisms] (P : C) (hP : P theorem projective_of_map_projective (adj : F ⊣ G) [F.Full] [F.Faithful] (P : C) (hP : Projective (F.obj P)) : Projective P where factors f g _ := by - haveI := Adjunction.leftAdjointPreservesColimits.{0, 0} adj + haveI := Adjunction.leftAdjoint_preservesColimits.{0, 0} adj rcases (@hP).1 (F.map f) (F.map g) with ⟨f', hf'⟩ use adj.unit.app _ ≫ G.map f' ≫ (inv <| adj.unit.app _) exact F.map_injective (by simpa) @@ -211,7 +211,7 @@ def mapProjectivePresentation (adj : F ⊣ G) [G.PreservesEpimorphisms] (X : C) p := F.obj Y.p projective := adj.map_projective _ Y.projective f := F.map Y.f - epi := have := Adjunction.leftAdjointPreservesColimits.{0, 0} adj; inferInstance + epi := have := Adjunction.leftAdjoint_preservesColimits.{0, 0} adj; inferInstance end Adjunction diff --git a/Mathlib/CategoryTheory/Preadditive/Yoneda/Limits.lean b/Mathlib/CategoryTheory/Preadditive/Yoneda/Limits.lean index 8f0743dd4b1fd..b4c0266ceba89 100644 --- a/Mathlib/CategoryTheory/Preadditive/Yoneda/Limits.lean +++ b/Mathlib/CategoryTheory/Preadditive/Yoneda/Limits.lean @@ -30,21 +30,22 @@ namespace CategoryTheory variable {C : Type u} [Category.{v} C] [Preadditive C] -instance preservesLimitsPreadditiveYonedaObj (X : C) : PreservesLimits (preadditiveYonedaObj X) := +instance preservesLimits_preadditiveYonedaObj (X : C) : PreservesLimits (preadditiveYonedaObj X) := have : PreservesLimits (preadditiveYonedaObj X ⋙ forget _) := (inferInstance : PreservesLimits (yoneda.obj X)) - preservesLimitsOfReflectsOfPreserves _ (forget _) + preservesLimits_of_reflects_of_preserves _ (forget _) -instance preservesLimitsPreadditiveCoyonedaObj (X : Cᵒᵖ) : +instance preservesLimits_preadditiveCoyonedaObj (X : Cᵒᵖ) : PreservesLimits (preadditiveCoyonedaObj X) := have : PreservesLimits (preadditiveCoyonedaObj X ⋙ forget _) := (inferInstance : PreservesLimits (coyoneda.obj X)) - preservesLimitsOfReflectsOfPreserves _ (forget _) + preservesLimits_of_reflects_of_preserves _ (forget _) -instance PreservesLimitsPreadditiveYoneda.obj (X : C) : PreservesLimits (preadditiveYoneda.obj X) := +instance preservesLimits_preadditiveYoneda_obj (X : C) : + PreservesLimits (preadditiveYoneda.obj X) := show PreservesLimits (preadditiveYonedaObj X ⋙ forget₂ _ _) from inferInstance -instance PreservesLimitsPreadditiveCoyoneda.obj (X : Cᵒᵖ) : +instance preservesLimits_preadditiveCoyoneda_obj (X : Cᵒᵖ) : PreservesLimits (preadditiveCoyoneda.obj X) := show PreservesLimits (preadditiveCoyonedaObj X ⋙ forget₂ _ _) from inferInstance diff --git a/Mathlib/CategoryTheory/Sites/Abelian.lean b/Mathlib/CategoryTheory/Sites/Abelian.lean index a4d7e152b10a4..9c89454d03a37 100644 --- a/Mathlib/CategoryTheory/Sites/Abelian.lean +++ b/Mathlib/CategoryTheory/Sites/Abelian.lean @@ -37,7 +37,7 @@ instance sheafIsAbelian : Abelian (Sheaf J D) := let adj := sheafificationAdjunction J D abelianOfAdjunction _ _ (asIso adj.counit) adj -attribute [local instance] preservesBinaryBiproductsOfPreservesBinaryProducts +attribute [local instance] preservesBinaryBiproducts_of_preservesBinaryProducts instance presheafToSheaf_additive : (presheafToSheaf J D).Additive := (presheafToSheaf J D).additive_of_preservesBinaryBiproducts diff --git a/Mathlib/CategoryTheory/Sites/Coherent/ExtensiveSheaves.lean b/Mathlib/CategoryTheory/Sites/Coherent/ExtensiveSheaves.lean index c4d16b62e04d7..e21fd44b0c627 100644 --- a/Mathlib/CategoryTheory/Sites/Coherent/ExtensiveSheaves.lean +++ b/Mathlib/CategoryTheory/Sites/Coherent/ExtensiveSheaves.lean @@ -90,10 +90,10 @@ theorem Presieve.isSheaf_iff_preservesFiniteProducts (F : Cᵒᵖ ⥤ Type w) : inferInstanceAs <| ∀ (i : α), Mono (Sigma.ι Z i) let i : K ≅ Discrete.functor (fun i ↦ op (Z i)) := Discrete.natIsoFunctor let _ : PreservesLimit (Discrete.functor (fun i ↦ op (Z i))) F := - Presieve.preservesProductOfIsSheafFor F ?_ initialIsInitial _ (coproductIsCoproduct Z) + Presieve.preservesProduct_of_isSheafFor F ?_ initialIsInitial _ (coproductIsCoproduct Z) (FinitaryExtensive.isPullback_initial_to_sigma_ι Z) (hF (Presieve.ofArrows Z (fun i ↦ Sigma.ι Z i)) ?_) - · exact preservesLimitOfIsoDiagram F i.symm + · exact preservesLimit_of_iso_diagram F i.symm · apply hF refine ⟨Empty, inferInstance, Empty.elim, IsEmpty.elim inferInstance, rfl, ⟨default,?_, ?_⟩⟩ · ext b @@ -114,24 +114,23 @@ theorem Presieve.isSheaf_iff_preservesFiniteProducts (F : Cᵒᵖ ⥤ Type w) : A presheaf on a category which is `FinitaryExtensive` is a sheaf iff it preserves finite products. -/ theorem Presheaf.isSheaf_iff_preservesFiniteProducts (F : Cᵒᵖ ⥤ D) : - IsSheaf (extensiveTopology C) F ↔ Nonempty (PreservesFiniteProducts F) := by + IsSheaf (extensiveTopology C) F ↔ PreservesFiniteProducts F := by constructor · intro h rw [IsSheaf] at h - refine ⟨⟨fun J _ ↦ ⟨fun {K} ↦ ⟨fun {c} hc ↦ ?_⟩⟩⟩⟩ + refine ⟨fun J _ ↦ ⟨fun {K} ↦ ⟨fun {c} hc ↦ ?_⟩⟩⟩ + constructor apply coyonedaJointlyReflectsLimits intro ⟨E⟩ specialize h E rw [Presieve.isSheaf_iff_preservesFiniteProducts] at h have : PreservesLimit K (F.comp (coyoneda.obj ⟨E⟩)) := (h.some.preserves J).preservesLimit - change IsLimit ((F.comp (coyoneda.obj ⟨E⟩)).mapCone c) - apply this.preserves - exact hc - · intro ⟨_⟩ E + exact isLimitOfPreserves (F.comp (coyoneda.obj ⟨E⟩)) hc + · intro _ E rw [Presieve.isSheaf_iff_preservesFiniteProducts] exact ⟨inferInstance⟩ -noncomputable instance (F : Sheaf (extensiveTopology C) D) : PreservesFiniteProducts F.val := - ((Presheaf.isSheaf_iff_preservesFiniteProducts F.val).mp F.cond).some +instance (F : Sheaf (extensiveTopology C) D) : PreservesFiniteProducts F.val := + (Presheaf.isSheaf_iff_preservesFiniteProducts F.val).mp F.cond end CategoryTheory diff --git a/Mathlib/CategoryTheory/Sites/Coherent/LocallySurjective.lean b/Mathlib/CategoryTheory/Sites/Coherent/LocallySurjective.lean index 3a79842c88fc9..bd98aeb91cd38 100644 --- a/Mathlib/CategoryTheory/Sites/Coherent/LocallySurjective.lean +++ b/Mathlib/CategoryTheory/Sites/Coherent/LocallySurjective.lean @@ -118,7 +118,7 @@ lemma regularTopology.isLocallySurjective_sheafOfTypes [Preregular C] [FinitaryP let i' : ((a : α) → (F.obj ⟨Z a⟩)) ≅ (F.obj ⟨∐ Z⟩) := (Types.productIso _).symm ≪≫ (PreservesProduct.iso F _).symm ≪≫ F.mapIso (opCoproductIsoProduct _).symm refine ⟨∐ Z, Sigma.desc π, inferInstance, i'.hom x, ?_⟩ - have := preservesLimitsOfShapeOfEquiv (Discrete.opposite α).symm G + have := preservesLimitsOfShape_of_equiv (Discrete.opposite α).symm G apply Concrete.isLimit_ext _ (isLimitOfPreserves G (coproductIsCoproduct Z).op) intro ⟨⟨a⟩⟩ simp only [Functor.comp_obj, Functor.op_obj, Discrete.functor_obj, Functor.mapCone_pt, diff --git a/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean b/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean index a20deb1ceff5d..e78d9e4dc5cb6 100644 --- a/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean +++ b/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean @@ -239,7 +239,7 @@ theorem isSheaf_coherent_iff_regular_and_extensive [Preregular C] [FinitaryPreEx theorem isSheaf_iff_preservesFiniteProducts_and_equalizerCondition [Preregular C] [FinitaryExtensive C] [h : ∀ {Y X : C} (f : Y ⟶ X) [EffectiveEpi f], HasPullback f f] : - IsSheaf (coherentTopology C) F ↔ Nonempty (PreservesFiniteProducts F) ∧ + IsSheaf (coherentTopology C) F ↔ PreservesFiniteProducts F ∧ EqualizerCondition F := by rw [isSheaf_coherent_iff_regular_and_extensive] exact and_congr (isSheaf_iff_preservesFiniteProducts _) @@ -247,12 +247,12 @@ theorem isSheaf_iff_preservesFiniteProducts_and_equalizerCondition noncomputable instance [Preregular C] [FinitaryExtensive C] (F : Sheaf (coherentTopology C) A) : PreservesFiniteProducts F.val := - ((Presheaf.isSheaf_iff_preservesFiniteProducts F.val).1 - ((Presheaf.isSheaf_coherent_iff_regular_and_extensive F.val).mp F.cond).1).some + (Presheaf.isSheaf_iff_preservesFiniteProducts F.val).1 + ((Presheaf.isSheaf_coherent_iff_regular_and_extensive F.val).mp F.cond).1 theorem isSheaf_iff_preservesFiniteProducts_of_projective [Preregular C] [FinitaryExtensive C] [∀ (X : C), Projective X] : - IsSheaf (coherentTopology C) F ↔ Nonempty (PreservesFiniteProducts F) := by + IsSheaf (coherentTopology C) F ↔ PreservesFiniteProducts F := by rw [isSheaf_coherent_iff_regular_and_extensive, and_iff_left (isSheaf_of_projective F), isSheaf_iff_preservesFiniteProducts] @@ -284,8 +284,8 @@ lemma isSheaf_coherent_of_hasPullbacks_comp [Preregular C] [FinitaryExtensive C] [h : ∀ {Y X : C} (f : Y ⟶ X) [EffectiveEpi f], HasPullback f f] [PreservesFiniteLimits s] (hF : IsSheaf (coherentTopology C) F) : IsSheaf (coherentTopology C) (F ⋙ s) := by rw [isSheaf_iff_preservesFiniteProducts_and_equalizerCondition (h := h)] at hF ⊢ - have := hF.1.some - refine ⟨⟨inferInstance⟩, fun _ _ π _ c hc ↦ ⟨?_⟩⟩ + have := hF.1 + refine ⟨inferInstance, fun _ _ π _ c hc ↦ ⟨?_⟩⟩ exact isLimitForkMapOfIsLimit s _ (hF.2 π c hc).some lemma isSheaf_coherent_of_hasPullbacks_of_comp [Preregular C] [FinitaryExtensive C] @@ -293,24 +293,24 @@ lemma isSheaf_coherent_of_hasPullbacks_of_comp [Preregular C] [FinitaryExtensive [ReflectsFiniteLimits s] (hF : IsSheaf (coherentTopology C) (F ⋙ s)) : IsSheaf (coherentTopology C) F := by rw [isSheaf_iff_preservesFiniteProducts_and_equalizerCondition (h := h)] at hF ⊢ - refine ⟨⟨⟨fun J _ ↦ ⟨fun {K} ↦ ⟨fun {c} hc ↦ ?_⟩⟩⟩⟩, fun _ _ π _ c hc ↦ ⟨?_⟩⟩ - · exact isLimitOfReflects s ((hF.1.some.1 J).1.1 hc) - · exact isLimitOfIsLimitForkMap s _ (hF.2 π c hc).some + obtain ⟨_, hF₂⟩ := hF + refine ⟨⟨fun J _ ↦ ⟨fun {K} ↦ ⟨fun {c} hc ↦ ?_⟩⟩⟩, fun _ _ π _ c hc ↦ ⟨?_⟩⟩ + · exact ⟨isLimitOfReflects s (isLimitOfPreserves (F ⋙ s) hc)⟩ + · exact isLimitOfIsLimitForkMap s _ (hF₂ π c hc).some lemma isSheaf_coherent_of_projective_comp [Preregular C] [FinitaryExtensive C] [∀ (X : C), Projective X] [PreservesFiniteProducts s] (hF : IsSheaf (coherentTopology C) F) : IsSheaf (coherentTopology C) (F ⋙ s) := by rw [isSheaf_iff_preservesFiniteProducts_of_projective] at hF ⊢ - have := hF.some - exact ⟨inferInstance⟩ + infer_instance lemma isSheaf_coherent_of_projective_of_comp [Preregular C] [FinitaryExtensive C] [∀ (X : C), Projective X] [ReflectsFiniteProducts s] (hF : IsSheaf (coherentTopology C) (F ⋙ s)) : IsSheaf (coherentTopology C) F := by rw [isSheaf_iff_preservesFiniteProducts_of_projective] at hF ⊢ - refine ⟨⟨fun J _ ↦ ⟨fun {K} ↦ ⟨fun {c} hc ↦ ?_⟩⟩⟩⟩ - exact isLimitOfReflects s ((hF.some.1 J).1.1 hc) + exact ⟨fun J _ ↦ ⟨fun {K} ↦ ⟨fun {c} hc ↦ + ⟨isLimitOfReflects s (isLimitOfPreserves (F ⋙ s) hc)⟩⟩⟩⟩ instance [Preregular C] [FinitaryExtensive C] [h : ∀ {Y X : C} (f : Y ⟶ X) [EffectiveEpi f], HasPullback f f] diff --git a/Mathlib/CategoryTheory/Sites/Equivalence.lean b/Mathlib/CategoryTheory/Sites/Equivalence.lean index de7ec32fc5e16..6953348d2088a 100644 --- a/Mathlib/CategoryTheory/Sites/Equivalence.lean +++ b/Mathlib/CategoryTheory/Sites/Equivalence.lean @@ -146,7 +146,7 @@ def transportSheafificationAdjunction : transportAndSheafify J K e A ⊣ sheafTo (transportIsoSheafToPresheaf _ _ _ _) noncomputable instance : PreservesFiniteLimits <| transportAndSheafify J K e A where - preservesFiniteLimits _ := compPreservesLimitsOfShape _ _ + preservesFiniteLimits _ := comp_preservesLimitsOfShape _ _ include K e in /-- Transport `HasSheafify` along an equivalence of sites. -/ diff --git a/Mathlib/CategoryTheory/Sites/LeftExact.lean b/Mathlib/CategoryTheory/Sites/LeftExact.lean index d12346e4674e9..4e6d5411bd86f 100644 --- a/Mathlib/CategoryTheory/Sites/LeftExact.lean +++ b/Mathlib/CategoryTheory/Sites/LeftExact.lean @@ -44,55 +44,61 @@ def coneCompEvaluationOfConeCompDiagramFunctorCompEvaluation {X : C} {K : Type m dsimp [diagramNatTrans] simp only [Multiequalizer.lift_ι, Category.assoc] } +/-- Auxiliary definition for `liftToDiagramLimitObj`. -/ +def liftToDiagramLimitObjAux {X : C} {K : Type max v u} [SmallCategory K] [HasLimitsOfShape K D] + {W : (J.Cover X)ᵒᵖ} (F : K ⥤ Cᵒᵖ ⥤ D) + (E : Cone (F ⋙ J.diagramFunctor D X ⋙ (evaluation (J.Cover X)ᵒᵖ D).obj W)) + (i : (unop W).Arrow) : + E.pt ⟶ (limit F).obj (op i.Y) := + (isLimitOfPreserves ((evaluation Cᵒᵖ D).obj (op i.Y)) (limit.isLimit F)).lift + (coneCompEvaluationOfConeCompDiagramFunctorCompEvaluation.{w, v, u} i E) + +@[reassoc (attr := simp)] +lemma liftToDiagramLimitObjAux_fac {X : C} {K : Type max v u} [SmallCategory K] + [HasLimitsOfShape K D] {W : (J.Cover X)ᵒᵖ} (F : K ⥤ Cᵒᵖ ⥤ D) + (E : Cone (F ⋙ J.diagramFunctor D X ⋙ (evaluation (J.Cover X)ᵒᵖ D).obj W)) + (i : (unop W).Arrow) (k : K) : + liftToDiagramLimitObjAux F E i ≫ (limit.π F k).app (op i.Y) = E.π.app k ≫ + Multiequalizer.ι ((unop W).index (F.obj k)) i := + IsLimit.fac _ _ _ + /-- An auxiliary definition to be used in the proof of the fact that `J.diagramFunctor D X` preserves limits. -/ abbrev liftToDiagramLimitObj {X : C} {K : Type max v u} [SmallCategory K] [HasLimitsOfShape K D] {W : (J.Cover X)ᵒᵖ} (F : K ⥤ Cᵒᵖ ⥤ D) (E : Cone (F ⋙ J.diagramFunctor D X ⋙ (evaluation (J.Cover X)ᵒᵖ D).obj W)) : E.pt ⟶ (J.diagram (limit F) X).obj W := - Multiequalizer.lift ((unop W).index (limit F)) E.pt - (fun i => (isLimitOfPreserves ((evaluation Cᵒᵖ D).obj (op i.Y)) (limit.isLimit F)).lift - (coneCompEvaluationOfConeCompDiagramFunctorCompEvaluation.{w, v, u} i E)) + Multiequalizer.lift ((unop W).index (limit F)) E.pt (liftToDiagramLimitObjAux F E) (by intro i - change (_ ≫ _) ≫ _ = (_ ≫ _) ≫ _ - dsimp [evaluateCombinedCones] - erw [Category.comp_id, Category.comp_id, Category.assoc, Category.assoc, ← - (limit.lift F _).naturality, ← (limit.lift F _).naturality, ← Category.assoc, ← - Category.assoc] - congr 1 - refine limit.hom_ext (fun j => ?_) - erw [Category.assoc, Category.assoc, limit.lift_π, limit.lift_π, limit.lift_π_assoc, - limit.lift_π_assoc, Category.assoc, Category.assoc, Multiequalizer.condition] + dsimp + ext k + dsimp + simp only [Category.assoc, NatTrans.naturality, liftToDiagramLimitObjAux_fac_assoc] + erw [Multiequalizer.condition] rfl) instance preservesLimit_diagramFunctor (X : C) (K : Type max v u) [SmallCategory K] [HasLimitsOfShape K D] (F : K ⥤ Cᵒᵖ ⥤ D) : PreservesLimit F (J.diagramFunctor D X) := - preservesLimitOfEvaluation _ _ fun W => - preservesLimitOfPreservesLimitCone (limit.isLimit _) + preservesLimit_of_evaluation _ _ fun W => + preservesLimit_of_preserves_limit_cone (limit.isLimit _) { lift := fun E => liftToDiagramLimitObj.{w, v, u} F E fac := by intro E k dsimp [diagramNatTrans] refine Multiequalizer.hom_ext _ _ _ (fun a => ?_) - simp only [Multiequalizer.lift_ι, Multiequalizer.lift_ι_assoc, Category.assoc] - change (_ ≫ _) ≫ _ = _ - dsimp [evaluateCombinedCones] - erw [Category.comp_id, Category.assoc, ← NatTrans.comp_app, limit.lift_π, limit.lift_π] - rfl + simp only [Multiequalizer.lift_ι, Multiequalizer.lift_ι_assoc, Category.assoc, + liftToDiagramLimitObjAux_fac] uniq := by intro E m hm refine Multiequalizer.hom_ext _ _ _ (fun a => limit_obj_ext (fun j => ?_)) - delta liftToDiagramLimitObj - erw [Multiequalizer.lift_ι, Category.assoc] - change _ = (_ ≫ _) ≫ _ - dsimp [evaluateCombinedCones] - erw [Category.comp_id, Category.assoc, ← NatTrans.comp_app, limit.lift_π, limit.lift_π] + dsimp [liftToDiagramLimitObj] + rw [Multiequalizer.lift_ι, Category.assoc, liftToDiagramLimitObjAux_fac, ← hm, + Category.assoc] dsimp - rw [← hm] - dsimp [diagramNatTrans] - simp } + rw [limit.lift_π] + dsimp } instance preservesLimitsOfShape_diagramFunctor (X : C) (K : Type max v u) [SmallCategory K] [HasLimitsOfShape K D] : @@ -169,8 +175,8 @@ instance preservesLimitsOfShape_plusFunctor (K : Type max v u) [SmallCategory K] [FinCategory K] [HasLimitsOfShape K D] [PreservesLimitsOfShape K (forget D)] [ReflectsLimitsOfShape K (forget D)] : PreservesLimitsOfShape K (J.plusFunctor D) := by - constructor; intro F; apply preservesLimitOfEvaluation; intro X - apply preservesLimitOfPreservesLimitCone (limit.isLimit F) + constructor; intro F; apply preservesLimit_of_evaluation; intro X + apply preservesLimit_of_preserves_limit_cone (limit.isLimit F) refine ⟨fun S => liftToPlusObjLimitObj.{w, v, u} F X.unop S, ?_, ?_⟩ · intro S k apply liftToPlusObjLimitObj_fac @@ -196,21 +202,21 @@ instance preservesLimitsOfShape_plusFunctor instance preserveFiniteLimits_plusFunctor [HasFiniteLimits D] [PreservesFiniteLimits (forget D)] [(forget D).ReflectsIsomorphisms] : PreservesFiniteLimits (J.plusFunctor D) := by - apply preservesFiniteLimitsOfPreservesFiniteLimitsOfSize.{max v u} + apply preservesFiniteLimits_of_preservesFiniteLimitsOfSize.{max v u} intro K _ _ - have : ReflectsLimitsOfShape K (forget D) := reflectsLimitsOfShapeOfReflectsIsomorphisms + have : ReflectsLimitsOfShape K (forget D) := reflectsLimitsOfShape_of_reflectsIsomorphisms apply preservesLimitsOfShape_plusFunctor.{w, v, u} instance preservesLimitsOfShape_sheafification (K : Type max v u) [SmallCategory K] [FinCategory K] [HasLimitsOfShape K D] [PreservesLimitsOfShape K (forget D)] [ReflectsLimitsOfShape K (forget D)] : PreservesLimitsOfShape K (J.sheafification D) := - Limits.compPreservesLimitsOfShape _ _ + Limits.comp_preservesLimitsOfShape _ _ instance preservesFiniteLimits_sheafification [HasFiniteLimits D] [PreservesFiniteLimits (forget D)] [(forget D).ReflectsIsomorphisms] : PreservesFiniteLimits (J.sheafification D) := - Limits.compPreservesFiniteLimits _ _ + Limits.comp_preservesFiniteLimits _ _ end CategoryTheory.GrothendieckTopology @@ -240,16 +246,16 @@ instance preservesLimitsOfShape_presheafToSheaf : · intro j j' show Fintype (ULift _) infer_instance - refine @preservesLimitsOfShapeOfEquiv _ _ _ _ _ _ _ _ e.symm _ (show _ from ?_) - constructor; intro F; constructor; intro S hS + refine @preservesLimitsOfShape_of_equiv _ _ _ _ _ _ _ _ e.symm _ (show _ from ?_) + constructor; intro F; constructor; intro S hS; constructor apply isLimitOfReflects (sheafToPresheaf J D) have : ReflectsLimitsOfShape (AsSmall.{max v u} (FinCategory.AsType K)) (forget D) := - reflectsLimitsOfShapeOfReflectsIsomorphisms + reflectsLimitsOfShape_of_reflectsIsomorphisms apply isLimitOfPreserves (J.sheafification D) hS instance preservesfiniteLimits_presheafToSheaf [HasFiniteLimits D] : PreservesFiniteLimits (plusPlusSheaf J D) := by - apply preservesFiniteLimitsOfPreservesFiniteLimitsOfSize.{max v u} + apply preservesFiniteLimits_of_preservesFiniteLimitsOfSize.{max v u} intros infer_instance diff --git a/Mathlib/CategoryTheory/Sites/Preserves.lean b/Mathlib/CategoryTheory/Sites/Preserves.lean index 233115606bae5..ffa55fc78899b 100644 --- a/Mathlib/CategoryTheory/Sites/Preserves.lean +++ b/Mathlib/CategoryTheory/Sites/Preserves.lean @@ -53,14 +53,15 @@ def isTerminal_of_isSheafFor_empty_presieve : IsTerminal (F.obj (op I)) := by choose t h using hF (by tauto) (by tauto) exact ⟨⟨fun _ ↦ t⟩, fun a ↦ by ext; exact h.2 _ (by tauto)⟩ +include hF in /-- If `F` is a presheaf which satisfies the sheaf condition with respect to the empty presieve on the initial object, then `F` preserves terminal objects. -/ -noncomputable -def preservesTerminalOfIsSheafForEmpty (hI : IsInitial I) : PreservesLimit (Functor.empty Cᵒᵖ) F := +lemma preservesTerminal_of_isSheaf_for_empty (hI : IsInitial I) : + PreservesLimit (Functor.empty.{0} Cᵒᵖ) F := have := hI.hasInitial - (preservesTerminalOfIso F + (preservesTerminal_of_iso F ((F.mapIso (terminalIsoIsTerminal (terminalOpOfInitial initialIsInitial)) ≪≫ (F.mapIso (initialIsoIsInitial hI).symm.op) ≪≫ (terminalIsoIsTerminal (isTerminal_of_isSheafFor_empty_presieve I F hF)).symm))) @@ -114,9 +115,8 @@ theorem isSheafFor_of_preservesProduct [PreservesLimit (Discrete.functor (fun x variable [HasInitial C] [∀ i, Mono (c.inj i)] (hd : Pairwise fun i j => IsPullback (initial.to _) (initial.to _) (c.inj i) (c.inj j)) -include hd -include hF hI in +include hd hF hI in /-- The two parallel maps in the equalizer diagram for the sheaf condition corresponding to the inclusion maps in a disjoint coproduct are equal. @@ -129,34 +129,33 @@ theorem firstMap_eq_secondMap : Equalizer.Presieve.Arrows.secondMap] by_cases hi : i = j · rw [hi, Mono.right_cancellation _ _ pullback.condition] - · have := preservesTerminalOfIsSheafForEmpty F hF hI + · have := preservesTerminal_of_isSheaf_for_empty F hF hI apply_fun (F.mapIso ((hd hi).isoPullback).op ≪≫ F.mapIso (terminalIsoIsTerminal (terminalOpOfInitial initialIsInitial)).symm ≪≫ (PreservesTerminal.iso F)).hom using injective_of_mono _ ext ⟨i⟩ exact i.elim +include hc hd hF hI in /-- If `F` is a presheaf which `IsSheafFor` a presieve of arrows and the empty presieve, then it preserves the product corresponding to the presieve of arrows. -/ -noncomputable -def preservesProductOfIsSheafFor - (hd : Pairwise fun i j => IsPullback (initial.to _) (initial.to _) (c.inj i) (c.inj j)) +lemma preservesProduct_of_isSheafFor (hF' : (ofArrows X c.inj).IsSheafFor F) : PreservesLimit (Discrete.functor (fun x ↦ op (X x))) F := by have : HasCoproduct X := ⟨⟨c, hc⟩⟩ - refine @PreservesProduct.ofIsoComparison _ _ _ _ F _ (fun x ↦ op (X x)) _ _ ?_ + refine @PreservesProduct.of_iso_comparison _ _ _ _ F _ (fun x ↦ op (X x)) _ _ ?_ rw [piComparison_fac (hc := hc)] refine IsIso.comp_isIso' inferInstance ?_ rw [isIso_iff_bijective, Function.bijective_iff_existsUnique] rw [Equalizer.Presieve.Arrows.sheaf_condition, Limits.Types.type_equalizer_iff_unique] at hF' exact fun b ↦ hF' b (congr_fun (firstMap_eq_secondMap F hF hI c hd) b) -include hF hI hc in +include hc hd hF hI in theorem isSheafFor_iff_preservesProduct : (ofArrows X c.inj).IsSheafFor F ↔ Nonempty (PreservesLimit (Discrete.functor (fun x ↦ op (X x))) F) := by - refine ⟨fun hF' ↦ ⟨preservesProductOfIsSheafFor _ hF hI c hc hd hF'⟩, fun hF' ↦ ?_⟩ + refine ⟨fun hF' ↦ ⟨preservesProduct_of_isSheafFor _ hF hI c hc hd hF'⟩, fun hF' ↦ ?_⟩ let _ := hF'.some exact isSheafFor_of_preservesProduct F c hc diff --git a/Mathlib/CategoryTheory/Sites/Pullback.lean b/Mathlib/CategoryTheory/Sites/Pullback.lean index 335206084917e..ccf9e1aa3527c 100644 --- a/Mathlib/CategoryTheory/Sites/Pullback.lean +++ b/Mathlib/CategoryTheory/Sites/Pullback.lean @@ -41,7 +41,7 @@ variable (J : GrothendieckTopology C) (K : GrothendieckTopology D) variable [ConcreteCategory.{v₁} A] [PreservesLimits (forget A)] [HasColimits A] [HasLimits A] variable [PreservesFilteredColimits (forget A)] [(forget A).ReflectsIsomorphisms] -attribute [local instance] reflectsLimitsOfReflectsIsomorphisms +attribute [local instance] reflectsLimits_of_reflectsIsomorphisms instance {X : C} : IsCofiltered (J.Cover X) := inferInstance @@ -54,8 +54,8 @@ def Functor.sheafPullback : Sheaf J A ⥤ Sheaf K A := instance [RepresentablyFlat G] : PreservesFiniteLimits (G.sheafPullback A J K) := by have : PreservesFiniteLimits (G.op.lan ⋙ presheafToSheaf K A) := - compPreservesFiniteLimits _ _ - apply compPreservesFiniteLimits + comp_preservesFiniteLimits _ _ + apply comp_preservesFiniteLimits /-- The pullback functor is left adjoint to the pushforward functor. -/ def Functor.sheafAdjunctionContinuous [Functor.IsContinuous.{v₁} G J K] : diff --git a/Mathlib/CategoryTheory/Sites/Sheaf.lean b/Mathlib/CategoryTheory/Sites/Sheaf.lean index e691068952e92..9ef621b55f223 100644 --- a/Mathlib/CategoryTheory/Sites/Sheaf.lean +++ b/Mathlib/CategoryTheory/Sites/Sheaf.lean @@ -670,7 +670,7 @@ theorem isSheaf_iff_isSheaf' : IsSheaf J P' ↔ IsSheaf' J P' := by rw [Equalizer.Presieve.sheaf_condition] refine ⟨?_⟩ refine isSheafForIsSheafFor' _ _ _ _ ?_ - letI := preservesSmallestLimitsOfPreservesLimits (coyoneda.obj (op U)) + letI := preservesSmallestLimits_of_preservesLimits (coyoneda.obj (op U)) apply isLimitOfPreserves apply Classical.choice (h _ S.arrows _) simpa @@ -692,7 +692,7 @@ theorem isSheaf_comp_of_isSheaf (s : A ⥤ B) [PreservesLimitsOfSize.{v₁, max theorem isSheaf_iff_isSheaf_comp (s : A ⥤ B) [HasLimitsOfSize.{v₁, max v₁ u₁} A] [PreservesLimitsOfSize.{v₁, max v₁ u₁} s] [s.ReflectsIsomorphisms] : IsSheaf J P ↔ IsSheaf J (P ⋙ s) := by - letI : ReflectsLimitsOfSize s := reflectsLimitsOfReflectsIsomorphisms + letI : ReflectsLimitsOfSize s := reflectsLimits_of_reflectsIsomorphisms exact ⟨isSheaf_comp_of_isSheaf J P s, isSheaf_of_isSheaf_comp J P s⟩ /-- @@ -707,7 +707,7 @@ hold. theorem isSheaf_iff_isSheaf_forget (s : A' ⥤ Type max v₁ u₁) [HasLimits A'] [PreservesLimits s] [s.ReflectsIsomorphisms] : IsSheaf J P' ↔ IsSheaf J (P' ⋙ s) := by have : HasLimitsOfSize.{v₁, max v₁ u₁} A' := hasLimitsOfSizeShrink.{_, _, u₁, 0} A' - have : PreservesLimitsOfSize.{v₁, max v₁ u₁} s := preservesLimitsOfSizeShrink.{_, 0, _, u₁} s + have : PreservesLimitsOfSize.{v₁, max v₁ u₁} s := preservesLimitsOfSize_shrink.{_, 0, _, u₁} s apply isSheaf_iff_isSheaf_comp end Concrete diff --git a/Mathlib/CategoryTheory/Sites/Sheafification.lean b/Mathlib/CategoryTheory/Sites/Sheafification.lean index 409605999eeb7..d8da476e5182c 100644 --- a/Mathlib/CategoryTheory/Sites/Sheafification.lean +++ b/Mathlib/CategoryTheory/Sites/Sheafification.lean @@ -55,7 +55,7 @@ theorem HasSheafify.mk' {F : (Cᵒᵖ ⥤ A) ⥤ Sheaf J A} (adj : F ⊣ sheafTo isRightAdjoint := ⟨F, ⟨adj⟩⟩ isLeftExact := ⟨by have : (sheafToPresheaf J A).IsRightAdjoint := ⟨_, ⟨adj⟩⟩ - exact ⟨fun _ _ _ ↦ preservesLimitsOfShapeOfNatIso + exact ⟨fun _ _ _ ↦ preservesLimitsOfShape_of_natIso (adj.leftAdjointUniq (Adjunction.ofIsRightAdjoint (sheafToPresheaf J A)))⟩⟩ /-- The sheafification functor, left adjoint to the inclusion. -/ diff --git a/Mathlib/CategoryTheory/Triangulated/Functor.lean b/Mathlib/CategoryTheory/Triangulated/Functor.lean index 9837ef05680cf..884a5b5ce77dc 100644 --- a/Mathlib/CategoryTheory/Triangulated/Functor.lean +++ b/Mathlib/CategoryTheory/Triangulated/Functor.lean @@ -173,8 +173,8 @@ instance (priority := 100) [F.IsTriangulated] : PreservesZeroMorphisms F where noncomputable instance [F.IsTriangulated] : PreservesLimitsOfShape (Discrete WalkingPair) F := by suffices ∀ (X₁ X₃ : C), IsIso (prodComparison F X₁ X₃) by - have := fun (X₁ X₃ : C) ↦ PreservesLimitPair.ofIsoProdComparison F X₁ X₃ - exact ⟨fun {K} ↦ preservesLimitOfIsoDiagram F (diagramIsoPair K).symm⟩ + have := fun (X₁ X₃ : C) ↦ PreservesLimitPair.of_iso_prod_comparison F X₁ X₃ + exact ⟨fun {K} ↦ preservesLimit_of_iso_diagram F (diagramIsoPair K).symm⟩ intro X₁ X₃ let φ : F.mapTriangle.obj (binaryProductTriangle X₁ X₃) ⟶ binaryProductTriangle (F.obj X₁) (F.obj X₃) := diff --git a/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean b/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean index fa43e61adcc47..5c1911f19648b 100644 --- a/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean +++ b/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean @@ -122,7 +122,7 @@ lemma mem_homologicalKernel_iff [F.IsHomological] [F.ShiftSequence ℤ] (X : C) noncomputable instance (priority := 100) [F.IsHomological] : PreservesLimitsOfShape (Discrete WalkingPair) F := by suffices ∀ (X₁ X₂ : C), PreservesLimit (pair X₁ X₂) F from - ⟨fun {X} => preservesLimitOfIsoDiagram F (diagramIsoPair X).symm⟩ + ⟨fun {X} => preservesLimit_of_iso_diagram F (diagramIsoPair X).symm⟩ intro X₁ X₂ have : HasBinaryBiproduct (F.obj X₁) (F.obj X₂) := HasBinaryBiproducts.has_binary_biproduct _ _ have : Mono (F.biprodComparison X₁ X₂) := by @@ -137,8 +137,8 @@ noncomputable instance (priority := 100) [F.IsHomological] : simp only [assoc, biprodComparison_fst, zero_comp, ← F.map_comp, biprod.inl_fst, F.map_id, comp_id] at hf rw [hf, zero_comp] - have : PreservesBinaryBiproduct X₁ X₂ F := preservesBinaryBiproductOfMonoBiprodComparison _ - apply Limits.preservesBinaryProductOfPreservesBinaryBiproduct + have : PreservesBinaryBiproduct X₁ X₂ F := preservesBinaryBiproduct_of_mono_biprodComparison _ + apply Limits.preservesBinaryProduct_of_preservesBinaryBiproduct instance (priority := 100) [F.IsHomological] : F.Additive := F.additive_of_preserves_binary_products diff --git a/Mathlib/Condensed/Discrete/Characterization.lean b/Mathlib/Condensed/Discrete/Characterization.lean index 04c5cbbfc2c17..842f3c46d0442 100644 --- a/Mathlib/Condensed/Discrete/Characterization.lean +++ b/Mathlib/Condensed/Discrete/Characterization.lean @@ -147,13 +147,13 @@ theorem isDiscrete_tfae (M : CondensedMod.{u} R) : rw [isDiscrete_iff_isDiscrete_forget, ((CondensedSet.isDiscrete_tfae _).out 0 6:)] intro S letI : PreservesFilteredColimitsOfSize.{u, u} (forget (ModuleCat R)) := - preservesFilteredColimitsOfSizeShrink.{u, u+1, u, u+1} _ + preservesFilteredColimitsOfSize_shrink.{u, u+1, u, u+1} _ exact ⟨isColimitOfPreserves (forget (ModuleCat R)) (h S).some⟩ tfae_have 1 → 7 := by intro h S rw [isDiscrete_iff_isDiscrete_forget, ((CondensedSet.isDiscrete_tfae _).out 0 6:)] at h letI : ReflectsFilteredColimitsOfSize.{u, u} (forget (ModuleCat R)) := - reflectsFilteredColimitsOfSizeShrink.{u, u+1, u, u+1} _ + reflectsFilteredColimitsOfSize_shrink.{u, u+1, u, u+1} _ exact ⟨isColimitOfReflects (forget (ModuleCat R)) (h S).some⟩ tfae_finish @@ -251,13 +251,13 @@ theorem isDiscrete_tfae (M : LightCondMod.{u} R) : rw [isDiscrete_iff_isDiscrete_forget, ((LightCondSet.isDiscrete_tfae _).out 0 5:)] intro S letI : PreservesFilteredColimitsOfSize.{0, 0} (forget (ModuleCat R)) := - preservesFilteredColimitsOfSizeShrink.{0, u, 0, u} _ + preservesFilteredColimitsOfSize_shrink.{0, u, 0, u} _ exact ⟨isColimitOfPreserves (forget (ModuleCat R)) (h S).some⟩ tfae_have 1 → 6 := by intro h S rw [isDiscrete_iff_isDiscrete_forget, ((LightCondSet.isDiscrete_tfae _).out 0 5:)] at h letI : ReflectsFilteredColimitsOfSize.{0, 0} (forget (ModuleCat R)) := - reflectsFilteredColimitsOfSizeShrink.{0, u, 0, u} _ + reflectsFilteredColimitsOfSize_shrink.{0, u, 0, u} _ exact ⟨isColimitOfReflects (forget (ModuleCat R)) (h S).some⟩ tfae_finish diff --git a/Mathlib/Condensed/Discrete/Colimit.lean b/Mathlib/Condensed/Discrete/Colimit.lean index b5f5decb0e322..b991d411c5c9a 100644 --- a/Mathlib/Condensed/Discrete/Colimit.lean +++ b/Mathlib/Condensed/Discrete/Colimit.lean @@ -202,7 +202,7 @@ noncomputable instance (X : Profinite) [Finite X] : let X' := (Countable.toSmall.{0} X).equiv_small.choose let e : X ≃ X' := (Countable.toSmall X).equiv_small.choose_spec.some have : Finite X' := .of_equiv X e - preservesLimitsOfShapeOfEquiv (Discrete.equivalence e.symm) F + preservesLimitsOfShape_of_equiv (Discrete.equivalence e.symm) F /-- Auxiliary definition for `isoFinYoneda`. -/ def isoFinYonedaComponents (X : Profinite.{u}) [Finite X] : @@ -473,7 +473,7 @@ noncomputable instance (X : FintypeCat.{u}) : PreservesLimitsOfShape (Discrete X let X' := (Countable.toSmall.{0} X).equiv_small.choose let e : X ≃ X' := (Countable.toSmall X).equiv_small.choose_spec.some have : Fintype X' := Fintype.ofEquiv X e - preservesLimitsOfShapeOfEquiv (Discrete.equivalence e.symm) F + preservesLimitsOfShape_of_equiv (Discrete.equivalence e.symm) F /-- Auxiliary definition for `isoFinYoneda`. -/ def isoFinYonedaComponents (X : LightProfinite.{u}) [Finite X] : diff --git a/Mathlib/Condensed/Explicit.lean b/Mathlib/Condensed/Explicit.lean index 541c0747e15fc..2efdfa8599be7 100644 --- a/Mathlib/Condensed/Explicit.lean +++ b/Mathlib/Condensed/Explicit.lean @@ -46,7 +46,7 @@ noncomputable def ofSheafStonean val := F cond := by rw [isSheaf_iff_preservesFiniteProducts_of_projective F] - exact ⟨⟨fun _ _ ↦ inferInstance⟩⟩ } + exact ⟨fun _ _ ↦ inferInstance⟩ } /-- The condensed object associated to a presheaf on `Stonean` whose postcomposition with the @@ -62,7 +62,7 @@ noncomputable def ofSheafForgetStonean cond := by apply isSheaf_coherent_of_projective_of_comp F (CategoryTheory.forget A) rw [isSheaf_iff_preservesFiniteProducts_of_projective] - exact ⟨⟨fun _ _ ↦ inferInstance⟩⟩ } + exact ⟨fun _ _ ↦ inferInstance⟩ } /-- The condensed object associated to a presheaf on `Profinite` which preserves finite products and @@ -76,7 +76,7 @@ noncomputable def ofSheafProfinite val := F cond := by rw [isSheaf_iff_preservesFiniteProducts_and_equalizerCondition F] - exact ⟨⟨⟨fun _ _ ↦ inferInstance⟩⟩, hF⟩ } + exact ⟨⟨fun _ _ ↦ inferInstance⟩, hF⟩ } /-- The condensed object associated to a presheaf on `Profinite` whose postcomposition with the @@ -93,7 +93,7 @@ noncomputable def ofSheafForgetProfinite cond := by apply isSheaf_coherent_of_hasPullbacks_of_comp F (CategoryTheory.forget A) rw [isSheaf_iff_preservesFiniteProducts_and_equalizerCondition] - exact ⟨⟨⟨fun _ _ ↦ inferInstance⟩⟩, hF⟩ } + exact ⟨⟨fun _ _ ↦ inferInstance⟩, hF⟩ } /-- The condensed object associated to a presheaf on `CompHaus` which preserves finite products and @@ -105,7 +105,7 @@ noncomputable def ofSheafCompHaus val := F cond := by rw [isSheaf_iff_preservesFiniteProducts_and_equalizerCondition F] - exact ⟨⟨⟨fun _ _ ↦ inferInstance⟩⟩, hF⟩ + exact ⟨⟨fun _ _ ↦ inferInstance⟩, hF⟩ /-- The condensed object associated to a presheaf on `CompHaus` whose postcomposition with the @@ -119,7 +119,7 @@ noncomputable def ofSheafForgetCompHaus cond := by apply isSheaf_coherent_of_hasPullbacks_of_comp F (CategoryTheory.forget A) rw [isSheaf_iff_preservesFiniteProducts_and_equalizerCondition] - exact ⟨⟨⟨fun _ _ ↦ inferInstance⟩⟩, hF⟩ + exact ⟨⟨fun _ _ ↦ inferInstance⟩, hF⟩ /-- A condensed object satisfies the equalizer condition. -/ theorem equalizerCondition (X : Condensed A) : EqualizerCondition X.val := @@ -128,13 +128,13 @@ theorem equalizerCondition (X : Condensed A) : EqualizerCondition X.val := /-- A condensed object preserves finite products. -/ noncomputable instance (X : Condensed A) : PreservesFiniteProducts X.val := isSheaf_iff_preservesFiniteProducts_and_equalizerCondition X.val |>.mp - X.cond |>.1.some + X.cond |>.1 /-- A condensed object regarded as a sheaf on `Profinite` preserves finite products. -/ noncomputable instance (X : Sheaf (coherentTopology Profinite.{u}) A) : PreservesFiniteProducts X.val := isSheaf_iff_preservesFiniteProducts_and_equalizerCondition X.val |>.mp - X.cond |>.1.some + X.cond |>.1 /-- A condensed object regarded as a sheaf on `Profinite` satisfies the equalizer condition. -/ theorem equalizerCondition_profinite (X : Sheaf (coherentTopology Profinite.{u}) A) : @@ -144,7 +144,7 @@ theorem equalizerCondition_profinite (X : Sheaf (coherentTopology Profinite.{u}) /-- A condensed object regarded as a sheaf on `Stonean` preserves finite products. -/ noncomputable instance (X : Sheaf (coherentTopology Stonean.{u}) A) : PreservesFiniteProducts X.val := - isSheaf_iff_preservesFiniteProducts_of_projective X.val |>.mp X.cond |>.some + isSheaf_iff_preservesFiniteProducts_of_projective X.val |>.mp X.cond end Condensed diff --git a/Mathlib/Condensed/Light/Explicit.lean b/Mathlib/Condensed/Light/Explicit.lean index 837d7f33f7fab..1c5823f86695f 100644 --- a/Mathlib/Condensed/Light/Explicit.lean +++ b/Mathlib/Condensed/Light/Explicit.lean @@ -42,7 +42,7 @@ noncomputable def ofSheafLightProfinite (F : LightProfinite.{u}ᵒᵖ ⥤ A) [Pr val := F cond := by rw [isSheaf_iff_preservesFiniteProducts_and_equalizerCondition F] - exact ⟨⟨⟨fun _ _ ↦ inferInstance⟩⟩, hF⟩ + exact ⟨⟨fun _ _ ↦ inferInstance⟩, hF⟩ /-- The light condensed object associated to a presheaf on `LightProfinite` whose postcomposition with @@ -57,7 +57,7 @@ noncomputable def ofSheafForgetLightProfinite cond := by apply isSheaf_coherent_of_hasPullbacks_of_comp F (CategoryTheory.forget A) rw [isSheaf_iff_preservesFiniteProducts_and_equalizerCondition] - exact ⟨⟨⟨fun _ _ ↦ inferInstance⟩⟩, hF⟩ + exact ⟨⟨fun _ _ ↦ inferInstance⟩, hF⟩ /-- A light condensed object satisfies the equalizer condition. -/ theorem equalizerCondition (X : LightCondensed A) : EqualizerCondition X.val := @@ -65,7 +65,7 @@ theorem equalizerCondition (X : LightCondensed A) : EqualizerCondition X.val := /-- A light condensed object preserves finite products. -/ noncomputable instance (X : LightCondensed A) : PreservesFiniteProducts X.val := - isSheaf_iff_preservesFiniteProducts_and_equalizerCondition X.val |>.mp X.cond |>.1.some + isSheaf_iff_preservesFiniteProducts_and_equalizerCondition X.val |>.mp X.cond |>.1 end LightCondensed diff --git a/Mathlib/Condensed/TopComparison.lean b/Mathlib/Condensed/TopComparison.lean index 92c2ed342f48f..aea4064452d23 100644 --- a/Mathlib/Condensed/TopComparison.lean +++ b/Mathlib/Condensed/TopComparison.lean @@ -89,8 +89,8 @@ the extensive topology. -/ noncomputable instance [PreservesFiniteCoproducts G] : PreservesFiniteProducts (yonedaPresheaf G X) := - have := preservesFiniteProductsOp G - ⟨fun _ ↦ compPreservesLimitsOfShape G.op (yonedaPresheaf' X)⟩ + have := preservesFiniteProducts_op G + ⟨fun _ ↦ comp_preservesLimitsOfShape G.op (yonedaPresheaf' X)⟩ section @@ -109,7 +109,7 @@ def TopCat.toSheafCompHausLike : cond := by have := CompHausLike.preregular hs rw [Presheaf.isSheaf_iff_preservesFiniteProducts_and_equalizerCondition] - refine ⟨⟨inferInstance⟩, ?_⟩ + refine ⟨inferInstance, ?_⟩ apply (config := { allowSynthFailures := true }) equalizerCondition_yonedaPresheaf (CompHausLike.compHausLikeToTop.{u} P) X intro Z B π he diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean index ac53d680e623b..e8f22b8f1f79b 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean @@ -119,7 +119,7 @@ instance : HasCoproducts.{u} LocallyRingedSpace.{u} := fun _ => noncomputable instance (J : Type _) : PreservesColimitsOfShape (Discrete.{u} J) forgetToSheafedSpace.{u} := ⟨fun {G} => - preservesColimitOfPreservesColimitCocone (coproductCofanIsColimit G) + preservesColimit_of_preserves_colimit_cocone (coproductCofanIsColimit G) ((colimit.isColimit (C := SheafedSpace.{u+1, u, u} CommRingCatMax.{u, u}) _).ofIsoColimit (Cocones.ext (Iso.refl _) fun _ => Category.comp_id _))⟩ @@ -320,8 +320,8 @@ noncomputable instance preservesCoequalizer : -- of colimit is provided later suffices PreservesColimit (parallelPair (F.map WalkingParallelPairHom.left) (F.map WalkingParallelPairHom.right)) forgetToSheafedSpace from - preservesColimitOfIsoDiagram _ (diagramIsoParallelPair F).symm - apply preservesColimitOfPreservesColimitCocone (coequalizerCoforkIsColimit _ _) + preservesColimit_of_iso_diagram _ (diagramIsoParallelPair F).symm + apply preservesColimit_of_preserves_colimit_cocone (coequalizerCoforkIsColimit _ _) apply (isColimitMapCoconeCoforkEquiv _ _).symm _ dsimp only [forgetToSheafedSpace] exact coequalizerIsCoequalizer _ _⟩ @@ -333,7 +333,7 @@ instance : HasColimits LocallyRingedSpace := noncomputable instance preservesColimits_forgetToSheafedSpace : PreservesColimits LocallyRingedSpace.forgetToSheafedSpace.{u} := - preservesColimitsOfPreservesCoequalizersAndCoproducts _ + preservesColimits_of_preservesCoequalizers_and_coproducts _ end LocallyRingedSpace diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index b937022700288..d2c8e5bec1786 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -456,8 +456,8 @@ instance pullbackToBaseIsOpenImmersion [IsOpenImmersion g] : rw [← limit.w (cospan f g) WalkingCospan.Hom.inl, cospan_map_inl] infer_instance -instance forgetPreservesLimitsOfLeft : PreservesLimit (cospan f g) (forget C) := - preservesLimitOfPreservesLimitCone (pullbackConeOfLeftIsLimit f g) +instance forget_preservesLimitsOfLeft : PreservesLimit (cospan f g) (forget C) := + preservesLimit_of_preserves_limit_cone (pullbackConeOfLeftIsLimit f g) (by apply (IsLimit.postcomposeHomEquiv (diagramIsoCospan _) _).toFun refine (IsLimit.equivIsoLimit ?_).toFun (limit.isLimit (cospan f.base g.base)) @@ -471,8 +471,8 @@ instance forgetPreservesLimitsOfLeft : PreservesLimit (cospan f g) (forget C) := · exact Category.comp_id _ · exact Category.comp_id _) -instance forgetPreservesLimitsOfRight : PreservesLimit (cospan g f) (forget C) := - preservesPullbackSymmetry (forget C) f g +instance forget_preservesLimitsOfRight : PreservesLimit (cospan g f) (forget C) := + preservesPullback_symmetry (forget C) f g theorem pullback_snd_isIso_of_range_subset (H : Set.range g.base ⊆ Set.range f.base) : IsIso (pullback.snd f g) := by @@ -665,18 +665,20 @@ instance forgetCreatesPullbackOfRight : CreatesLimit (cospan g f) forget := (eqToIso (show pullback _ _ = pullback _ _ by congr) ≪≫ HasLimit.isoOfNatIso (diagramIsoCospan _).symm) -instance sheafedSpaceForgetPreservesOfLeft : PreservesLimit (cospan f g) (SheafedSpace.forget C) := - @Limits.compPreservesLimit _ _ _ _ _ _ (cospan f g) _ _ forget (PresheafedSpace.forget C) +instance sheafedSpace_forgetPreserves_of_left : + PreservesLimit (cospan f g) (SheafedSpace.forget C) := + @Limits.comp_preservesLimit _ _ _ _ _ _ (cospan f g) _ _ forget (PresheafedSpace.forget C) inferInstance <| by have : PreservesLimit (cospan ((cospan f g ⋙ forget).map Hom.inl) ((cospan f g ⋙ forget).map Hom.inr)) (PresheafedSpace.forget C) := by dsimp infer_instance - apply preservesLimitOfIsoDiagram _ (diagramIsoCospan _).symm + apply preservesLimit_of_iso_diagram _ (diagramIsoCospan _).symm -instance sheafedSpaceForgetPreservesOfRight : PreservesLimit (cospan g f) (SheafedSpace.forget C) := - preservesPullbackSymmetry _ _ _ +instance sheafedSpace_forgetPreserves_of_right : + PreservesLimit (cospan g f) (SheafedSpace.forget C) := + preservesPullback_symmetry _ _ _ instance sheafedSpace_hasPullback_of_left : HasPullback f g := hasLimit_of_created (cospan f g) forget @@ -1021,17 +1023,17 @@ instance pullback_to_base_isOpenImmersion [LocallyRingedSpace.IsOpenImmersion g] rw [← limit.w (cospan f g) WalkingCospan.Hom.inl, cospan_map_inl] infer_instance -instance forgetPreservesPullbackOfLeft : +instance forget_preservesPullbackOfLeft : PreservesLimit (cospan f g) LocallyRingedSpace.forgetToSheafedSpace := - preservesLimitOfPreservesLimitCone (pullbackConeOfLeftIsLimit f g) <| by + preservesLimit_of_preserves_limit_cone (pullbackConeOfLeftIsLimit f g) <| by apply (isLimitMapConePullbackConeEquiv _ _).symm.toFun apply isLimitOfIsLimitPullbackConeMap SheafedSpace.forgetToPresheafedSpace exact PresheafedSpace.IsOpenImmersion.pullbackConeOfLeftIsLimit f.1 g.1 -instance forgetToPresheafedSpacePreservesPullbackOfLeft : +instance forgetToPresheafedSpace_preservesPullback_of_left : PreservesLimit (cospan f g) (LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forgetToPresheafedSpace) := - preservesLimitOfPreservesLimitCone (pullbackConeOfLeftIsLimit f g) <| by + preservesLimit_of_preserves_limit_cone (pullbackConeOfLeftIsLimit f g) <| by apply (isLimitMapConePullbackConeEquiv _ _).symm.toFun exact PresheafedSpace.IsOpenImmersion.pullbackConeOfLeftIsLimit f.1 g.1 @@ -1040,7 +1042,7 @@ instance forgetToPresheafedSpacePreservesOpenImmersion : ((LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forgetToPresheafedSpace).map f) := H -instance forgetToTopPreservesPullbackOfLeft : +instance forgetToTop_preservesPullback_of_left : PreservesLimit (cospan f g) (LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forget _) := by change PreservesLimit _ <| @@ -1056,35 +1058,35 @@ instance forgetToTopPreservesPullbackOfLeft : dsimp; infer_instance have : PreservesLimit (cospan f g ⋙ forgetToSheafedSpace ⋙ SheafedSpace.forgetToPresheafedSpace) (PresheafedSpace.forget CommRingCat) := by - apply preservesLimitOfIsoDiagram _ (diagramIsoCospan _).symm - apply Limits.compPreservesLimit + apply preservesLimit_of_iso_diagram _ (diagramIsoCospan _).symm + apply Limits.comp_preservesLimit -instance forgetReflectsPullbackOfLeft : +instance forget_reflectsPullback_of_left : ReflectsLimit (cospan f g) LocallyRingedSpace.forgetToSheafedSpace := - reflectsLimitOfReflectsIsomorphisms _ _ + reflectsLimit_of_reflectsIsomorphisms _ _ -instance forgetPreservesPullbackOfRight : +instance forget_preservesPullback_of_right : PreservesLimit (cospan g f) LocallyRingedSpace.forgetToSheafedSpace := - preservesPullbackSymmetry _ _ _ + preservesPullback_symmetry _ _ _ -instance forgetToPresheafedSpacePreservesPullbackOfRight : +instance forgetToPresheafedSpace_preservesPullback_of_right : PreservesLimit (cospan g f) (LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forgetToPresheafedSpace) := - preservesPullbackSymmetry _ _ _ + preservesPullback_symmetry _ _ _ -instance forgetReflectsPullbackOfRight : +instance forget_reflectsPullback_of_right : ReflectsLimit (cospan g f) LocallyRingedSpace.forgetToSheafedSpace := - reflectsLimitOfReflectsIsomorphisms _ _ + reflectsLimit_of_reflectsIsomorphisms _ _ -instance forgetToPresheafedSpaceReflectsPullbackOfLeft : +instance forgetToPresheafedSpace_reflectsPullback_of_left : ReflectsLimit (cospan f g) (LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forgetToPresheafedSpace) := - reflectsLimitOfReflectsIsomorphisms _ _ + reflectsLimit_of_reflectsIsomorphisms _ _ -instance forgetToPresheafedSpaceReflectsPullbackOfRight : +instance forgetToPresheafedSpace_reflectsPullback_of_right : ReflectsLimit (cospan g f) (LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forgetToPresheafedSpace) := - reflectsLimitOfReflectsIsomorphisms _ _ + reflectsLimit_of_reflectsIsomorphisms _ _ theorem pullback_snd_isIso_of_range_subset (H' : Set.range g.base ⊆ Set.range f.base) : IsIso (pullback.snd f g) := by diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean index e178d24170737..4d7c3d0c180bc 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean @@ -263,7 +263,7 @@ instance : HasColimitsOfShape J (PresheafedSpace.{_, _, v} C) where has_colimit F := ⟨colimitCocone F, colimitCoconeIsColimit F⟩ instance : PreservesColimitsOfShape J (PresheafedSpace.forget.{u, v, v} C) := - ⟨fun {F} => preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit F) <| by + ⟨fun {F} => preservesColimit_of_preserves_colimit_cocone (colimitCoconeIsColimit F) <| by apply IsColimit.ofIsoColimit (colimit.isColimit _) fapply Cocones.ext · rfl @@ -278,15 +278,12 @@ instance instHasColimits [HasLimits C] : HasColimits (PresheafedSpace.{_, _, v} /-- The underlying topological space of a colimit of presheaved spaces is the colimit of the underlying topological spaces. -/ -instance forgetPreservesColimits [HasLimits C] : PreservesColimits (PresheafedSpace.forget C) where +instance forget_preservesColimits [HasLimits C] : + PreservesColimits (PresheafedSpace.forget.{_, _, v} C) where preservesColimitsOfShape {J 𝒥} := - { preservesColimit := fun {F} => - preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit F) - (by apply IsColimit.ofIsoColimit (colimit.isColimit _) - fapply Cocones.ext - · rfl - · intro j - simp) } + { preservesColimit := fun {F} => preservesColimit_of_preserves_colimit_cocone + (colimitCoconeIsColimit F) + (IsColimit.ofIsoColimit (colimit.isColimit _) (Cocones.ext (Iso.refl _))) } /-- The components of the colimit of a diagram of `PresheafedSpace C` is obtained via taking componentwise limits. diff --git a/Mathlib/Geometry/RingedSpace/SheafedSpace.lean b/Mathlib/Geometry/RingedSpace/SheafedSpace.lean index eace03ab08bbd..1941ecd8cc471 100644 --- a/Mathlib/Geometry/RingedSpace/SheafedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/SheafedSpace.lean @@ -206,8 +206,8 @@ noncomputable instance [HasLimits C] : instance [HasLimits C] : HasColimits.{v} (SheafedSpace C) := hasColimits_of_hasColimits_createsColimits forgetToPresheafedSpace -noncomputable instance [HasLimits C] : PreservesColimits (forget C) := - Limits.compPreservesColimits forgetToPresheafedSpace (PresheafedSpace.forget C) +noncomputable instance [HasLimits C] : PreservesColimits (forget.{_, _, v} C) := + Limits.comp_preservesColimits forgetToPresheafedSpace (PresheafedSpace.forget C) section ConcreteCategory diff --git a/Mathlib/RepresentationTheory/Action/Basic.lean b/Mathlib/RepresentationTheory/Action/Basic.lean index 2a8960d7da208..7c61dd710f8eb 100644 --- a/Mathlib/RepresentationTheory/Action/Basic.lean +++ b/Mathlib/RepresentationTheory/Action/Basic.lean @@ -265,13 +265,13 @@ def functorCategoryEquivalenceCompEvaluation : (functorCategoryEquivalence V G).functor ⋙ (evaluation _ _).obj PUnit.unit ≅ forget V G := Iso.refl _ -noncomputable instance instPreservesLimitsForget [HasLimits V] : - Limits.PreservesLimits (forget V G) := - Limits.preservesLimitsOfNatIso (Action.functorCategoryEquivalenceCompEvaluation V G) +noncomputable instance preservesLimits_forget [HasLimits V] : + PreservesLimits (forget V G) := + Limits.preservesLimits_of_natIso (Action.functorCategoryEquivalenceCompEvaluation V G) -noncomputable instance instPreservesColimitsForget [HasColimits V] : +noncomputable instance preservesColimits_forget [HasColimits V] : PreservesColimits (forget V G) := - preservesColimitsOfNatIso (Action.functorCategoryEquivalenceCompEvaluation V G) + preservesColimits_of_natIso (Action.functorCategoryEquivalenceCompEvaluation V G) -- TODO construct categorical images? end Forget diff --git a/Mathlib/RepresentationTheory/Action/Limits.lean b/Mathlib/RepresentationTheory/Action/Limits.lean index 4e1d7d252a087..104f4244fde0e 100644 --- a/Mathlib/RepresentationTheory/Action/Limits.lean +++ b/Mathlib/RepresentationTheory/Action/Limits.lean @@ -69,80 +69,80 @@ variable {C : Type t₁} [Category.{t₂} C] /-- `F : C ⥤ SingleObj G ⥤ V` preserves the limit of some `K : J ⥤ C` if it does evaluated at `SingleObj.star G`. -/ -private def SingleObj.preservesLimit (F : C ⥤ SingleObj G ⥤ V) +private lemma SingleObj.preservesLimit (F : C ⥤ SingleObj G ⥤ V) {J : Type w₁} [Category.{w₂} J] (K : J ⥤ C) (h : PreservesLimit K (F ⋙ (evaluation (SingleObj G) V).obj (SingleObj.star G))) : PreservesLimit K F := by - apply preservesLimitOfEvaluation + apply preservesLimit_of_evaluation intro _ exact h /-- `F : C ⥤ Action V G` preserves the limit of some `K : J ⥤ C` if if it does after postcomposing with the forgetful functor `Action V G ⥤ V`. -/ -noncomputable def preservesLimitOfPreserves (F : C ⥤ Action V G) {J : Type w₁} +lemma preservesLimit_of_preserves (F : C ⥤ Action V G) {J : Type w₁} [Category.{w₂} J] (K : J ⥤ C) (h : PreservesLimit K (F ⋙ Action.forget V G)) : PreservesLimit K F := by let F' : C ⥤ SingleObj G ⥤ V := F ⋙ (Action.functorCategoryEquivalence V G).functor have : PreservesLimit K F' := SingleObj.preservesLimit _ _ h - apply preservesLimitOfReflectsOfPreserves F (Action.functorCategoryEquivalence V G).functor + apply preservesLimit_of_reflects_of_preserves F (Action.functorCategoryEquivalence V G).functor /-- `F : C ⥤ Action V G` preserves limits of some shape `J` if it does after postcomposing with the forgetful functor `Action V G ⥤ V`. -/ -noncomputable def preservesLimitsOfShapeOfPreserves (F : C ⥤ Action V G) {J : Type w₁} +lemma preservesLimitsOfShape_of_preserves (F : C ⥤ Action V G) {J : Type w₁} [Category.{w₂} J] (h : PreservesLimitsOfShape J (F ⋙ Action.forget V G)) : PreservesLimitsOfShape J F := by constructor intro K - apply Action.preservesLimitOfPreserves + apply Action.preservesLimit_of_preserves exact PreservesLimitsOfShape.preservesLimit /-- `F : C ⥤ Action V G` preserves limits of some size if it does after postcomposing with the forgetful functor `Action V G ⥤ V`. -/ -noncomputable def preservesLimitsOfSizeOfPreserves (F : C ⥤ Action V G) +lemma preservesLimitsOfSize_of_preserves (F : C ⥤ Action V G) (h : PreservesLimitsOfSize.{w₂, w₁} (F ⋙ Action.forget V G)) : PreservesLimitsOfSize.{w₂, w₁} F := by constructor intro J _ - apply Action.preservesLimitsOfShapeOfPreserves + apply Action.preservesLimitsOfShape_of_preserves exact PreservesLimitsOfSize.preservesLimitsOfShape /-- `F : C ⥤ SingleObj G ⥤ V` preserves the colimit of some `K : J ⥤ C` if it does evaluated at `SingleObj.star G`. -/ -private def SingleObj.preservesColimit (F : C ⥤ SingleObj G ⥤ V) +private lemma SingleObj.preservesColimit (F : C ⥤ SingleObj G ⥤ V) {J : Type w₁} [Category.{w₂} J] (K : J ⥤ C) (h : PreservesColimit K (F ⋙ (evaluation (SingleObj G) V).obj (SingleObj.star G))) : PreservesColimit K F := by - apply preservesColimitOfEvaluation + apply preservesColimit_of_evaluation intro _ exact h /-- `F : C ⥤ Action V G` preserves the colimit of some `K : J ⥤ C` if if it does after postcomposing with the forgetful functor `Action V G ⥤ V`. -/ -noncomputable def preservesColimitOfPreserves (F : C ⥤ Action V G) {J : Type w₁} +lemma preservesColimit_of_preserves (F : C ⥤ Action V G) {J : Type w₁} [Category.{w₂} J] (K : J ⥤ C) (h : PreservesColimit K (F ⋙ Action.forget V G)) : PreservesColimit K F := by let F' : C ⥤ SingleObj G ⥤ V := F ⋙ (Action.functorCategoryEquivalence V G).functor have : PreservesColimit K F' := SingleObj.preservesColimit _ _ h - apply preservesColimitOfReflectsOfPreserves F (Action.functorCategoryEquivalence V G).functor + apply preservesColimit_of_reflects_of_preserves F (Action.functorCategoryEquivalence V G).functor /-- `F : C ⥤ Action V G` preserves colimits of some shape `J` if it does after postcomposing with the forgetful functor `Action V G ⥤ V`. -/ -noncomputable def preservesColimitsOfShapeOfPreserves (F : C ⥤ Action V G) {J : Type w₁} +lemma preservesColimitsOfShape_of_preserves (F : C ⥤ Action V G) {J : Type w₁} [Category.{w₂} J] (h : PreservesColimitsOfShape J (F ⋙ Action.forget V G)) : PreservesColimitsOfShape J F := by constructor intro K - apply Action.preservesColimitOfPreserves + apply Action.preservesColimit_of_preserves exact PreservesColimitsOfShape.preservesColimit /-- `F : C ⥤ Action V G` preserves colimits of some size if it does after postcomposing with the forgetful functor `Action V G ⥤ V`. -/ -noncomputable def preservesColimitsOfSizeOfPreserves (F : C ⥤ Action V G) +lemma preservesColimitsOfSize_of_preserves (F : C ⥤ Action V G) (h : PreservesColimitsOfSize.{w₂, w₁} (F ⋙ Action.forget V G)) : PreservesColimitsOfSize.{w₂, w₁} F := by constructor intro J _ - apply Action.preservesColimitsOfShapeOfPreserves + apply Action.preservesColimitsOfShape_of_preserves exact PreservesColimitsOfSize.preservesColimitsOfShape end Preservation @@ -168,7 +168,7 @@ noncomputable instance [HasFiniteLimits V] : PreservesFiniteLimits (Action.forge constructor intro _ _ _ infer_instance - apply compPreservesFiniteLimits + apply comp_preservesFiniteLimits noncomputable instance [HasFiniteColimits V] : PreservesFiniteColimits (Action.forget V G) := by show PreservesFiniteColimits ((Action.functorCategoryEquivalence V G).functor ⋙ @@ -177,24 +177,24 @@ noncomputable instance [HasFiniteColimits V] : PreservesFiniteColimits (Action.f constructor intro _ _ _ infer_instance - apply compPreservesFiniteColimits + apply comp_preservesFiniteColimits -noncomputable instance {J : Type w₁} [Category.{w₂} J] (F : J ⥤ Action V G) : +instance {J : Type w₁} [Category.{w₂} J] (F : J ⥤ Action V G) : ReflectsLimit F (Action.forget V G) where - reflects h := by + reflects h := ⟨by apply isLimitOfReflects ((Action.functorCategoryEquivalence V G).functor) - exact evaluationJointlyReflectsLimits _ (fun _ => h) + exact evaluationJointlyReflectsLimits _ (fun _ => h)⟩ -noncomputable instance {J : Type w₁} [Category.{w₂} J] : +instance {J : Type w₁} [Category.{w₂} J] : ReflectsLimitsOfShape J (Action.forget V G) where -noncomputable instance : ReflectsLimits (Action.forget V G) where +instance : ReflectsLimits (Action.forget V G) where -noncomputable instance {J : Type w₁} [Category.{w₂} J] (F : J ⥤ Action V G) : +instance {J : Type w₁} [Category.{w₂} J] (F : J ⥤ Action V G) : ReflectsColimit F (Action.forget V G) where - reflects h := by + reflects h := ⟨by apply isColimitOfReflects ((Action.functorCategoryEquivalence V G).functor) - exact evaluationJointlyReflectsColimits _ (fun _ => h) + exact evaluationJointlyReflectsColimits _ (fun _ => h)⟩ noncomputable instance {J : Type w₁} [Category.{w₂} J] : ReflectsColimitsOfShape J (Action.forget V G) where diff --git a/Mathlib/RepresentationTheory/Rep.lean b/Mathlib/RepresentationTheory/Rep.lean index 2af958302e683..d4ddfa12b4c67 100644 --- a/Mathlib/RepresentationTheory/Rep.lean +++ b/Mathlib/RepresentationTheory/Rep.lean @@ -122,10 +122,10 @@ instance {V : Type u} [AddCommGroup V] [Module k V] (ρ : Representation k G V) -- Porting note: the two following instances were found automatically in mathlib3 noncomputable instance : PreservesLimits (forget₂ (Rep k G) (ModuleCat.{u} k)) := - Action.instPreservesLimitsForget.{u} _ _ + Action.preservesLimits_forget.{u} _ _ noncomputable instance : PreservesColimits (forget₂ (Rep k G) (ModuleCat.{u} k)) := - Action.instPreservesColimitsForget.{u} _ _ + Action.preservesColimits_forget.{u} _ _ /- Porting note: linter complains `simp` unfolds some types in the LHS, so have removed `@[simp]`. -/ diff --git a/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Basic.lean b/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Basic.lean index 803016fe59879..2e6184e78b0ae 100644 --- a/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Basic.lean +++ b/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Basic.lean @@ -253,7 +253,7 @@ end instance : Limits.PreservesLimits profiniteGrpToProfinite.{u} where preservesLimitsOfShape := { - preservesLimit := fun {F} ↦ CategoryTheory.Limits.preservesLimitOfPreservesLimitCone + preservesLimit := fun {F} ↦ CategoryTheory.Limits.preservesLimit_of_preserves_limit_cone (limitConeIsLimit F) (Profinite.limitConeIsLimit (F ⋙ profiniteGrpToProfinite)) } end ProfiniteGrp diff --git a/Mathlib/Topology/Category/CompHausLike/Limits.lean b/Mathlib/Topology/Category/CompHausLike/Limits.lean index 8f8cb1165e429..7739d95e40a78 100644 --- a/Mathlib/Topology/Category/CompHausLike/Limits.lean +++ b/Mathlib/Topology/Category/CompHausLike/Limits.lean @@ -175,8 +175,8 @@ instance (P) [HasExplicitFiniteCoproducts.{0} P] : PreservesFiniteCoproducts (compHausLikeToTop P) := by refine ⟨fun J hJ ↦ ⟨fun {F} ↦ ?_⟩⟩ suffices PreservesColimit (Discrete.functor (F.obj ∘ Discrete.mk)) (compHausLikeToTop P) from - preservesColimitOfIsoDiagram _ Discrete.natIsoFunctor.symm - apply preservesColimitOfPreservesColimitCocone (CompHausLike.finiteCoproduct.isColimit _) + preservesColimit_of_iso_diagram _ Discrete.natIsoFunctor.symm + apply preservesColimit_of_preserves_colimit_cocone (CompHausLike.finiteCoproduct.isColimit _) exact TopCat.sigmaCofanIsColimit _ /-- The functor to another `CompHausLike` preserves finite coproducts if they exist. -/ @@ -185,7 +185,7 @@ noncomputable instance {P' : TopCat.{u} → Prop} PreservesFiniteCoproducts (toCompHausLike h) := by have : PreservesFiniteCoproducts (toCompHausLike h ⋙ compHausLikeToTop P') := inferInstanceAs (PreservesFiniteCoproducts (compHausLikeToTop _)) - exact preservesFiniteCoproductsOfReflectsOfPreserves (toCompHausLike h) (compHausLikeToTop P') + exact preservesFiniteCoproducts_of_reflects_of_preserves (toCompHausLike h) (compHausLikeToTop P') end FiniteCoproducts @@ -290,7 +290,7 @@ noncomputable instance : CreatesLimit (cospan f g) (compHausLikeToTop P) := by /-- The functor to `TopCat` preserves pullbacks. -/ noncomputable instance : PreservesLimit (cospan f g) (compHausLikeToTop P) := - preservesLimitOfCreatesLimitAndHasLimit _ _ + preservesLimit_of_createsLimit_and_hasLimit _ _ /-- The functor to another `CompHausLike` preserves pullbacks. -/ noncomputable instance {P' : TopCat → Prop} @@ -298,7 +298,7 @@ noncomputable instance {P' : TopCat → Prop} PreservesLimit (cospan f g) (toCompHausLike h) := by have : PreservesLimit (cospan f g) (toCompHausLike h ⋙ compHausLikeToTop P') := inferInstanceAs (PreservesLimit _ (compHausLikeToTop _)) - exact preservesLimitOfReflectsOfPreserves (toCompHausLike h) (compHausLikeToTop P') + exact preservesLimit_of_reflects_of_preserves (toCompHausLike h) (compHausLikeToTop P') variable (P) in /-- diff --git a/Mathlib/Topology/Category/LightProfinite/Basic.lean b/Mathlib/Topology/Category/LightProfinite/Basic.lean index 3e60e9a5ca8ad..362a23b4df054 100644 --- a/Mathlib/Topology/Category/LightProfinite/Basic.lean +++ b/Mathlib/Topology/Category/LightProfinite/Basic.lean @@ -175,8 +175,8 @@ noncomputable instance createsCountableLimits {J : Type v} [SmallCategory J] [Co instance : HasCountableLimits LightProfinite where out _ := { has_limit := fun F ↦ ⟨limitCone F, limitConeIsLimit F⟩ } -noncomputable instance : PreservesLimitsOfShape ℕᵒᵖ (forget LightProfinite.{u}) := - have : PreservesLimitsOfSize.{0, 0} (forget Profinite.{u}) := preservesLimitsOfSizeShrink _ +instance : PreservesLimitsOfShape ℕᵒᵖ (forget LightProfinite.{u}) := + have : PreservesLimitsOfSize.{0, 0} (forget Profinite.{u}) := preservesLimitsOfSize_shrink _ inferInstanceAs (PreservesLimitsOfShape ℕᵒᵖ (lightToProfinite ⋙ forget Profinite)) variable {X Y : LightProfinite.{u}} (f : X ⟶ Y) diff --git a/Mathlib/Topology/Category/Profinite/Basic.lean b/Mathlib/Topology/Category/Profinite/Basic.lean index 503c4020cd505..e15cf996568bd 100644 --- a/Mathlib/Topology/Category/Profinite/Basic.lean +++ b/Mathlib/Topology/Category/Profinite/Basic.lean @@ -216,8 +216,8 @@ instance hasLimits : Limits.HasLimits Profinite := instance hasColimits : Limits.HasColimits Profinite := hasColimits_of_reflective profiniteToCompHaus -noncomputable instance forgetPreservesLimits : Limits.PreservesLimits (forget Profinite) := by - apply Limits.compPreservesLimits Profinite.toTopCat (forget TopCat) +instance forget_preservesLimits : Limits.PreservesLimits (forget Profinite) := by + apply Limits.comp_preservesLimits Profinite.toTopCat (forget TopCat) theorem epi_iff_surjective {X Y : Profinite.{u}} (f : X ⟶ Y) : Epi f ↔ Function.Surjective f := by constructor diff --git a/Mathlib/Topology/Category/Stonean/Adjunctions.lean b/Mathlib/Topology/Category/Stonean/Adjunctions.lean index dd99b65695aed..da6e4d47f184f 100644 --- a/Mathlib/Topology/Category/Stonean/Adjunctions.lean +++ b/Mathlib/Topology/Category/Stonean/Adjunctions.lean @@ -53,6 +53,6 @@ noncomputable def stoneCechAdjunction : typeToStonean ⊣ (forget Stonean) := /-- The forgetful functor from Stonean spaces, being a right adjoint, preserves limits. -/ noncomputable instance forget.preservesLimits : Limits.PreservesLimits (forget Stonean) := - rightAdjointPreservesLimits stoneCechAdjunction + rightAdjoint_preservesLimits stoneCechAdjunction end Stonean diff --git a/Mathlib/Topology/Category/TopCat/Limits/Basic.lean b/Mathlib/Topology/Category/TopCat/Limits/Basic.lean index e7833bab71d89..4ed86b7b0ecfd 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Basic.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Basic.lean @@ -113,14 +113,15 @@ instance topCat_hasLimitsOfSize : HasLimitsOfSize.{v, v} TopCat.{max v u} where instance topCat_hasLimits : HasLimits TopCat.{u} := TopCat.topCat_hasLimitsOfSize.{u, u} -instance forgetPreservesLimitsOfSize : PreservesLimitsOfSize forget where +instance forget_preservesLimitsOfSize : + PreservesLimitsOfSize.{v, v} (forget : TopCat.{max v u} ⥤ _) where preservesLimitsOfShape {_} := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v,u} F) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit.{v,u} F) (Types.limitConeIsLimit.{v,u} (F ⋙ forget)) } -instance forgetPreservesLimits : PreservesLimits forget := - TopCat.forgetPreservesLimitsOfSize.{u,u} +instance forget_preservesLimits : PreservesLimits (forget : TopCat.{u} ⥤ _) := + TopCat.forget_preservesLimitsOfSize.{u, u} /-- A choice of colimit cocone for a functor `F : J ⥤ TopCat`. Generally you should just use `colimit.cocone F`, unless you need the actual definition @@ -171,15 +172,15 @@ instance topCat_hasColimitsOfSize : HasColimitsOfSize.{v,v} TopCat.{max v u} whe instance topCat_hasColimits : HasColimits TopCat.{u} := TopCat.topCat_hasColimitsOfSize.{u, u} -instance forgetPreservesColimitsOfSize : - PreservesColimitsOfSize.{v, v} forget where +instance forget_preservesColimitsOfSize : + PreservesColimitsOfSize.{v, v} (forget : TopCat.{max u v} ⥤ _) where preservesColimitsOfShape := { preservesColimit := fun {F} => - preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit F) + preservesColimit_of_preserves_colimit_cocone (colimitCoconeIsColimit F) (Types.TypeMax.colimitCoconeIsColimit (F ⋙ forget)) } -instance forgetPreservesColimits : PreservesColimits (forget : TopCat.{u} ⥤ Type u) := - TopCat.forgetPreservesColimitsOfSize.{u, u} +instance forget_preservesColimits : PreservesColimits (forget : TopCat.{u} ⥤ Type u) := + TopCat.forget_preservesColimitsOfSize.{u, u} /-- The terminal object of `Top` is `PUnit`. -/ def isTerminalPUnit : IsTerminal (TopCat.of PUnit.{u + 1}) := diff --git a/Mathlib/Topology/Category/TopCat/Yoneda.lean b/Mathlib/Topology/Category/TopCat/Yoneda.lean index d898269d9ac23..f724695c46e33 100644 --- a/Mathlib/Topology/Category/TopCat/Yoneda.lean +++ b/Mathlib/Topology/Category/TopCat/Yoneda.lean @@ -62,9 +62,9 @@ noncomputable instance : PreservesFiniteProducts (yonedaPresheaf'.{w, w'} Y) whe preserves _ _ := { preservesLimit := fun {K} => have : ∀ {α : Type} (X : α → TopCat), PreservesLimit (Discrete.functor (fun x ↦ op (X x))) - (yonedaPresheaf'.{w, w'} Y) := fun X => @PreservesProduct.ofIsoComparison _ _ _ _ + (yonedaPresheaf'.{w, w'} Y) := fun X => @PreservesProduct.of_iso_comparison _ _ _ _ (yonedaPresheaf' Y) _ (fun x ↦ op (X x)) _ _ (by rw [piComparison_fac]; infer_instance) let i : K ≅ Discrete.functor (fun i ↦ op (unop (K.obj ⟨i⟩))) := Discrete.natIsoFunctor - preservesLimitOfIsoDiagram _ i.symm } + preservesLimit_of_iso_diagram _ i.symm } end ContinuousMap diff --git a/Mathlib/Topology/Sheaves/SheafCondition/PairwiseIntersections.lean b/Mathlib/Topology/Sheaves/SheafCondition/PairwiseIntersections.lean index 18d81939f34e9..196cd67b22551 100644 --- a/Mathlib/Topology/Sheaves/SheafCondition/PairwiseIntersections.lean +++ b/Mathlib/Topology/Sheaves/SheafCondition/PairwiseIntersections.lean @@ -273,10 +273,10 @@ theorem isSheaf_iff_isSheafPreservesLimitPairwiseIntersections : rw [isSheaf_iff_isSheafPairwiseIntersections] constructor · intro h ι U - exact ⟨preservesLimitOfPreservesLimitCone (Pairwise.coconeIsColimit U).op (h U).some⟩ + exact ⟨preservesLimit_of_preserves_limit_cone (Pairwise.coconeIsColimit U).op (h U).some⟩ · intro h ι U haveI := (h U).some - exact ⟨PreservesLimit.preserves (Pairwise.coconeIsColimit U).op⟩ + exact ⟨isLimitOfPreserves _ (Pairwise.coconeIsColimit U).op⟩ end TopCat.Presheaf From 1aefbadf18578666efa249758f32e80fe369e560 Mon Sep 17 00:00:00 2001 From: Pim Otte Date: Tue, 19 Nov 2024 08:43:58 +0000 Subject: [PATCH 100/829] feat(Combinatorics/SimpleGraph/Matching): add `IsMatching.exists_of_disjoint_sets_of_card_eq` (#18905) This lemma supports at least two use cases: Matchings in cliques and matchings with a (sub)set of universal vertices. In preparation for Tutte's theorem. Co-authored-by: Pim Otte --- .../Combinatorics/SimpleGraph/Matching.lean | 28 +++++++++++++++++++ 1 file changed, 28 insertions(+) diff --git a/Mathlib/Combinatorics/SimpleGraph/Matching.lean b/Mathlib/Combinatorics/SimpleGraph/Matching.lean index 14f2cbfc960a6..d52240d416332 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Matching.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Matching.lean @@ -134,6 +134,34 @@ lemma IsMatching.coeSubgraph {G' : Subgraph G} {M : Subgraph G'.coe} (hM : M.IsM · obtain ⟨_, hw', hvw⟩ := (coeSubgraph_adj _ _ _).mp hy rw [← hw.2 ⟨y, hw'⟩ hvw] +lemma IsMatching.exists_of_disjoint_sets_of_equiv {s t : Set V} (h : Disjoint s t) + (f : s ≃ t) (hadj : ∀ v : s, G.Adj v (f v)) : + ∃ M : Subgraph G, M.verts = s ∪ t ∧ M.IsMatching := by + use { + verts := s ∪ t + Adj := fun v w ↦ (∃ h : v ∈ s, f ⟨v, h⟩ = w) ∨ (∃ h : w ∈ s, f ⟨w, h⟩ = v) + adj_sub := by + intro v w h + obtain (⟨hv, rfl⟩ | ⟨hw, rfl⟩) := h + · exact hadj ⟨v, _⟩ + · exact (hadj ⟨w, _⟩).symm + edge_vert := by aesop } + + simp only [Subgraph.IsMatching, Set.mem_union, true_and] + intro v hv + cases' hv with hl hr + · use f ⟨v, hl⟩ + simp only [hl, exists_const, true_or, exists_true_left, true_and] + rintro y (rfl | ⟨hys, rfl⟩) + · rfl + · exact (h.ne_of_mem hl (f ⟨y, hys⟩).coe_prop rfl).elim + · use f.symm ⟨v, hr⟩ + simp only [Subtype.coe_eta, Equiv.apply_symm_apply, Subtype.coe_prop, exists_const, or_true, + true_and] + rintro y (⟨hy, rfl⟩ | ⟨hy, rfl⟩) + · exact (h.ne_of_mem hy hr rfl).elim + · simp + protected lemma IsMatching.map {G' : SimpleGraph W} {M : Subgraph G} (f : G →g G') (hf : Injective f) (hM : M.IsMatching) : (M.map f).IsMatching := by rintro _ ⟨v, hv, rfl⟩ From 8ad393d08a4074b0dff2cac9609c3bd7d78041f2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 19 Nov 2024 08:43:59 +0000 Subject: [PATCH 101/829] feat(SetTheory/ZFC/Basic): results on pairs (#19144) We add a few lemmas on unordered pairs and use them to golf down the injectivity of the Kuratowski pair. --- Mathlib/SetTheory/ZFC/Basic.lean | 49 +++++++++++++++++++------------- 1 file changed, 29 insertions(+), 20 deletions(-) diff --git a/Mathlib/SetTheory/ZFC/Basic.lean b/Mathlib/SetTheory/ZFC/Basic.lean index 8d884f87ac28e..c7a0c35cdfcef 100644 --- a/Mathlib/SetTheory/ZFC/Basic.lean +++ b/Mathlib/SetTheory/ZFC/Basic.lean @@ -892,6 +892,24 @@ theorem singleton_nonempty (u : ZFSet) : ZFSet.Nonempty {u} := theorem mem_pair {x y z : ZFSet.{u}} : x ∈ ({y, z} : ZFSet) ↔ x = y ∨ x = z := by simp +@[simp] +theorem pair_eq_singleton (x : ZFSet) : {x, x} = ({x} : ZFSet) := by + ext + simp + +@[simp] +theorem pair_eq_singleton_iff {x y z : ZFSet} : ({x, y} : ZFSet) = {z} ↔ x = z ∧ y = z := by + refine ⟨fun h ↦ ?_, ?_⟩ + · rw [← mem_singleton, ← mem_singleton] + simp [← h] + · rintro ⟨rfl, rfl⟩ + exact pair_eq_singleton y + +@[simp] +theorem singleton_eq_pair_iff {x y z : ZFSet} : ({x} : ZFSet) = {y, z} ↔ x = y ∧ x = z := by + rw [eq_comm, pair_eq_singleton_iff] + simp_rw [eq_comm] + /-- `omega` is the first infinite von Neumann ordinal -/ def omega : ZFSet := mk PSet.omega @@ -1194,7 +1212,7 @@ theorem toSet_pair (x y : ZFSet.{u}) : (pair x y).toSet = {{x}, {x, y}} := by si /-- A subset of pairs `{(a, b) ∈ x × y | p a b}` -/ def pairSep (p : ZFSet.{u} → ZFSet.{u} → Prop) (x y : ZFSet.{u}) : ZFSet.{u} := - ZFSet.sep (fun z => ∃ a ∈ x, ∃ b ∈ y, z = pair a b ∧ p a b) (powerset (powerset (x ∪ y))) + (powerset (powerset (x ∪ y))).sep fun z => ∃ a ∈ x, ∃ b ∈ y, z = pair a b ∧ p a b @[simp] theorem mem_pairSep {p} {x y z : ZFSet.{u}} : @@ -1207,26 +1225,17 @@ theorem mem_pairSep {p} {x y z : ZFSet.{u}} : exact Or.inl ax · rintro (rfl | rfl) <;> [left; right] <;> assumption -theorem pair_injective : Function.Injective2 pair := fun x x' y y' H => by - have ae := ZFSet.ext_iff.1 H - simp only [pair, mem_pair] at ae - obtain rfl : x = x' := by - cases' (ae {x}).1 (by simp) with h h - · exact singleton_injective h - · have m : x' ∈ ({x} : ZFSet) := by simp [h] - rw [mem_singleton.mp m] - have he : x = y → y = y' := by +theorem pair_injective : Function.Injective2 pair := by + intro x x' y y' H + simp_rw [ZFSet.ext_iff, pair, mem_pair] at H + obtain rfl : x = x' := And.left <| by simpa [or_and_left] using (H {x}).1 (Or.inl rfl) + have he : y = x → y = y' := by rintro rfl - cases' (ae {x, y'}).2 (by simp only [eq_self_iff_true, or_true]) with xy'x xy'xx - · rw [eq_comm, ← mem_singleton, ← xy'x, mem_pair] - exact Or.inr rfl - · simpa [eq_comm] using (ZFSet.ext_iff.1 xy'xx y').1 (by simp) - obtain xyx | xyy' := (ae {x, y}).1 (by simp) - · obtain rfl := mem_singleton.mp ((ZFSet.ext_iff.1 xyx y).1 <| by simp) - simp [he rfl] - · obtain rfl | yy' := mem_pair.mp ((ZFSet.ext_iff.1 xyy' y).1 <| by simp) - · simp [he rfl] - · simp [yy'] + simpa [eq_comm] using H {y, y'} + have hx := H {x, y} + simp_rw [pair_eq_singleton_iff, true_and, or_true, true_iff] at hx + refine ⟨rfl, hx.elim he fun hy ↦ Or.elim ?_ he id⟩ + simpa using ZFSet.ext_iff.1 hy y @[simp] theorem pair_inj {x y x' y' : ZFSet} : pair x y = pair x' y' ↔ x = x' ∧ y = y' := From a5783c915dea554ddb4c2234e50b25a709fb1777 Mon Sep 17 00:00:00 2001 From: Jiang Jiedong Date: Tue, 19 Nov 2024 08:56:06 +0000 Subject: [PATCH 102/829] feat(FieldTheory/Minpoly): minpoly K x splits implies minpoly K (algebraMap K L a - x) splits (#17369) Add `minpoly K x` splits implies `minpoly K (- x)` splits and `minpoly K (algebraMap K L a - x)` splits. This is after #17093 . Co-authored-by: Jiang Jiedong <107380768+jjdishere@users.noreply.github.com> --- Mathlib/Algebra/Polynomial/AlgebraMap.lean | 45 +++++++++++-- Mathlib/Algebra/Polynomial/Splits.lean | 76 +++++++++++++++++----- Mathlib/RingTheory/Adjoin/Field.lean | 12 ++++ 3 files changed, 112 insertions(+), 21 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/AlgebraMap.lean b/Mathlib/Algebra/Polynomial/AlgebraMap.lean index 6318eb32f3a95..3431db816584a 100644 --- a/Mathlib/Algebra/Polynomial/AlgebraMap.lean +++ b/Mathlib/Algebra/Polynomial/AlgebraMap.lean @@ -306,22 +306,46 @@ theorem algEquivOfCompEqX_eq_iff (p q p' q' : R[X]) theorem algEquivOfCompEqX_symm (p q : R[X]) (hpq : p.comp q = X) (hqp : q.comp p = X) : (algEquivOfCompEqX p q hpq hqp).symm = algEquivOfCompEqX q p hqp hpq := rfl +/-- The automorphism of the polynomial algebra given by `p(X) ↦ p(a * X + b)`, + with inverse `p(X) ↦ p(a⁻¹ * (X - b))`. -/ +@[simps!] +def algEquivCMulXAddC {R : Type*} [CommRing R] (a b : R) [Invertible a] : R[X] ≃ₐ[R] R[X] := + algEquivOfCompEqX (C a * X + C b) (C ⅟ a * (X - C b)) + (by simp [← C_mul, ← mul_assoc]) (by simp [← C_mul, ← mul_assoc]) + +theorem algEquivCMulXAddC_symm_eq {R : Type*} [CommRing R] (a b : R) [Invertible a] : + (algEquivCMulXAddC a b).symm = algEquivCMulXAddC (⅟ a) (- ⅟ a * b) := by + ext p : 1 + simp only [algEquivCMulXAddC_symm_apply, neg_mul, algEquivCMulXAddC_apply, map_neg, map_mul] + congr + simp [mul_add, sub_eq_add_neg] + /-- The automorphism of the polynomial algebra given by `p(X) ↦ p(X+t)`, with inverse `p(X) ↦ p(X-t)`. -/ @[simps!] -def algEquivAevalXAddC {R} [CommRing R] (t : R) : R[X] ≃ₐ[R] R[X] := +def algEquivAevalXAddC {R : Type*} [CommRing R] (t : R) : R[X] ≃ₐ[R] R[X] := algEquivOfCompEqX (X + C t) (X - C t) (by simp) (by simp) @[simp] -theorem algEquivAevalXAddC_eq_iff {R} [CommRing R] (t t' : R) : +theorem algEquivAevalXAddC_eq_iff {R : Type*} [CommRing R] (t t' : R) : algEquivAevalXAddC t = algEquivAevalXAddC t' ↔ t = t' := by simp [algEquivAevalXAddC] @[simp] -theorem algEquivAevalXAddC_symm {R} [CommRing R] (t : R) : +theorem algEquivAevalXAddC_symm {R : Type*} [CommRing R] (t : R) : (algEquivAevalXAddC t).symm = algEquivAevalXAddC (-t) := by simp [algEquivAevalXAddC, sub_eq_add_neg] +/-- The involutive automorphism of the polynomial algebra given by `p(X) ↦ p(-X)`. -/ +@[simps!] +def algEquivAevalNegX {R : Type*} [CommRing R] : R[X] ≃ₐ[R] R[X] := + algEquivOfCompEqX (-X) (-X) (by simp) (by simp) + +theorem comp_neg_X_comp_neg_X {R : Type*} [CommRing R] (p : R[X]) : + (p.comp (-X)).comp (-X) = p := by + rw [comp_assoc] + simp only [neg_comp, X_comp, neg_neg, comp_X] + theorem aeval_algHom (f : A →ₐ[R] B) (x : A) : aeval (f x) = f.comp (aeval x) := algHom_ext <| by simp only [aeval_X, AlgHom.comp_apply] @@ -570,16 +594,25 @@ lemma comp_X_add_C_eq_zero_iff : p.comp (X + C t) = 0 ↔ p = 0 := lemma comp_X_add_C_ne_zero_iff : p.comp (X + C t) ≠ 0 ↔ p ≠ 0 := comp_X_add_C_eq_zero_iff.not +lemma dvd_comp_C_mul_X_add_C_iff (p q : R[X]) (a b : R) [Invertible a] : + p ∣ q.comp (C a * X + C b) ↔ p.comp (C ⅟ a * (X - C b)) ∣ q := by + convert map_dvd_iff <| algEquivCMulXAddC a b using 2 + simp [← comp_eq_aeval, comp_assoc, ← mul_assoc, ← C_mul] + lemma dvd_comp_X_sub_C_iff (p q : R[X]) (a : R) : p ∣ q.comp (X - C a) ↔ p.comp (X + C a) ∣ q := by - convert (map_dvd_iff <| algEquivAevalXAddC a).symm using 2 - rw [C_eq_algebraMap, algEquivAevalXAddC_apply, ← comp_eq_aeval] - simp [comp_assoc] + let _ := invertibleOne (α := R) + simpa using dvd_comp_C_mul_X_add_C_iff p q 1 (-a) lemma dvd_comp_X_add_C_iff (p q : R[X]) (a : R) : p ∣ q.comp (X + C a) ↔ p.comp (X - C a) ∣ q := by simpa using dvd_comp_X_sub_C_iff p q (-a) +lemma dvd_comp_neg_X_iff (p q : R[X]) : p ∣ q.comp (-X) ↔ p.comp (-X) ∣ q := by + let _ := invertibleOne (α := R) + let _ := invertibleNeg (α := R) 1 + simpa using dvd_comp_C_mul_X_add_C_iff p q (-1) 0 + variable [IsDomain R] lemma units_coeff_zero_smul (c : R[X]ˣ) (p : R[X]) : (c : R[X]).coeff 0 • p = c * p := by diff --git a/Mathlib/Algebra/Polynomial/Splits.lean b/Mathlib/Algebra/Polynomial/Splits.lean index 98d5a93538fdf..918c15651b4a4 100644 --- a/Mathlib/Algebra/Polynomial/Splits.lean +++ b/Mathlib/Algebra/Polynomial/Splits.lean @@ -139,26 +139,72 @@ theorem splits_X_pow (n : ℕ) : (X ^ n).Splits i := theorem splits_id_iff_splits {f : K[X]} : (f.map i).Splits (RingHom.id L) ↔ f.Splits i := by rw [splits_map_iff, RingHom.id_comp] -theorem Splits.comp_X_sub_C {i : L →+* F} (a : L) {f : L[X]} - (h : f.Splits i) : (f.comp (X - C a)).Splits i := by +variable {i} + +theorem Splits.comp_of_map_degree_le_one {f : K[X]} {p : K[X]} (hd : (p.map i).degree ≤ 1) + (h : f.Splits i) : (f.comp p).Splits i := by + by_cases hzero : map i (f.comp p) = 0 + · exact Or.inl hzero cases h with | inl h0 => - left - simp only [map_eq_zero] at h0 ⊢ - exact h0.symm ▸ zero_comp + exact Or.inl <| map_comp i _ _ ▸ h0.symm ▸ zero_comp | inr h => right intro g irr dvd - rw [map_comp, Polynomial.map_sub, map_X, map_C, dvd_comp_X_sub_C_iff] at dvd - have := h (irr.map (algEquivAevalXAddC _)) dvd - rw [degree_eq_natDegree irr.ne_zero] - rwa [algEquivAevalXAddC_apply, ← comp_eq_aeval, - degree_eq_natDegree (fun h => WithBot.bot_ne_one (h ▸ this)), - natDegree_comp, natDegree_X_add_C, mul_one] at this - -theorem Splits.comp_X_add_C {i : L →+* F} (a : L) {f : L[X]} - (h : f.Splits i) : (f.comp (X + C a)).Splits i := by - simpa only [map_neg, sub_neg_eq_add] using h.comp_X_sub_C (-a) + rw [map_comp] at dvd hzero + cases lt_or_eq_of_le hd with + | inl hd => + rw [eq_C_of_degree_le_zero (Nat.WithBot.lt_one_iff_le_zero.mp hd), comp_C] at dvd hzero + refine False.elim (irr.1 (isUnit_of_dvd_unit dvd ?_)) + simpa using hzero + | inr hd => + let _ := invertibleOfNonzero (leadingCoeff_ne_zero.mpr + (ne_zero_of_degree_gt (n := ⊥) (by rw [hd]; decide))) + rw [eq_X_add_C_of_degree_eq_one hd, dvd_comp_C_mul_X_add_C_iff _ _] at dvd + have := h (irr.map (algEquivCMulXAddC _ _).symm) dvd + rw [degree_eq_natDegree irr.ne_zero] + rwa [algEquivCMulXAddC_symm_apply, ← comp_eq_aeval, + degree_eq_natDegree (fun h => WithBot.bot_ne_one (h ▸ this)), + natDegree_comp, natDegree_C_mul (invertibleInvOf.ne_zero), + natDegree_X_sub_C, mul_one] at this + +theorem splits_iff_comp_splits_of_degree_eq_one {f : K[X]} {p : K[X]} (hd : (p.map i).degree = 1) : + f.Splits i ↔ (f.comp p).Splits i := by + rw [← splits_id_iff_splits, ← splits_id_iff_splits (f := f.comp p), map_comp] + refine ⟨fun h => Splits.comp_of_map_degree_le_one + (le_of_eq (map_id (R := L) ▸ hd)) h, fun h => ?_⟩ + let _ := invertibleOfNonzero (leadingCoeff_ne_zero.mpr + (ne_zero_of_degree_gt (n := ⊥) (by rw [hd]; decide))) + have : (map i f) = ((map i f).comp (map i p)).comp ((C ⅟ (map i p).leadingCoeff * + (X - C ((map i p).coeff 0)))) := by + rw [comp_assoc] + nth_rw 1 [eq_X_add_C_of_degree_eq_one hd] + simp only [coeff_map, invOf_eq_inv, mul_sub, ← C_mul, add_comp, mul_comp, C_comp, X_comp, + ← mul_assoc] + simp + refine this ▸ Splits.comp_of_map_degree_le_one ?_ h + simp [degree_C (inv_ne_zero (Invertible.ne_zero (a := (map i p).leadingCoeff)))] + +/-- +This is a weaker variant of `Splits.comp_of_map_degree_le_one`, +but its conditions are easier to check. +-/ +theorem Splits.comp_of_degree_le_one {f : K[X]} {p : K[X]} (hd : p.degree ≤ 1) + (h : f.Splits i) : (f.comp p).Splits i := + Splits.comp_of_map_degree_le_one ((degree_map_le i _).trans hd) h + +theorem Splits.comp_X_sub_C (a : K) {f : K[X]} + (h : f.Splits i) : (f.comp (X - C a)).Splits i := + Splits.comp_of_degree_le_one (degree_X_sub_C_le _) h + +theorem Splits.comp_X_add_C (a : K) {f : K[X]} + (h : f.Splits i) : (f.comp (X + C a)).Splits i := + Splits.comp_of_degree_le_one (by simpa using degree_X_sub_C_le (-a)) h + +theorem Splits.comp_neg_X {f : K[X]} (h : f.Splits i) : (f.comp (-X)).Splits i := + Splits.comp_of_degree_le_one (by simpa using degree_X_sub_C_le (0 : K)) h + +variable (i) theorem exists_root_of_splits' {f : K[X]} (hs : Splits i f) (hf0 : degree (f.map i) ≠ 0) : ∃ x, eval₂ i x f = 0 := diff --git a/Mathlib/RingTheory/Adjoin/Field.lean b/Mathlib/RingTheory/Adjoin/Field.lean index 25bf90e6eb79c..d791b712416ee 100644 --- a/Mathlib/RingTheory/Adjoin/Field.lean +++ b/Mathlib/RingTheory/Adjoin/Field.lean @@ -97,6 +97,13 @@ theorem IsIntegral.mem_range_algebraMap_of_minpoly_splits [Algebra K L] [IsScala x ∈ (algebraMap K L).range := int.mem_range_algHom_of_minpoly_splits h (IsScalarTower.toAlgHom R K L) +theorem minpoly_neg_splits [Algebra K L] {x : L} (g : (minpoly K x).Splits (algebraMap K L)) : + (minpoly K (-x)).Splits (algebraMap K L) := by + rw [minpoly.neg] + apply splits_mul _ _ g.comp_neg_X + simpa only [map_pow, map_neg, map_one] using + splits_C (algebraMap K L) ((-1) ^ (minpoly K x).natDegree) + theorem minpoly_add_algebraMap_splits [Algebra K L] {x : L} (r : K) (g : (minpoly K x).Splits (algebraMap K L)) : (minpoly K (x + algebraMap K L r)).Splits (algebraMap K L) := by @@ -112,6 +119,11 @@ theorem minpoly_algebraMap_add_splits [Algebra K L] {x : L} (r : K) (minpoly K (algebraMap K L r + x)).Splits (algebraMap K L) := by simpa only [add_comm] using minpoly_add_algebraMap_splits r g +theorem minpoly_algebraMap_sub_splits [Algebra K L] {x : L} (r : K) + (g : (minpoly K x).Splits (algebraMap K L)) : + (minpoly K (algebraMap K L r - x)).Splits (algebraMap K L) := by + simpa only [neg_sub] using minpoly_neg_splits (minpoly_sub_algebraMap_splits r g) + end variable [Algebra K M] [IsScalarTower R K M] {x : M} From 08c691878eacbbe7c022e11194894f840318d608 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Tue, 19 Nov 2024 09:20:30 +0000 Subject: [PATCH 103/829] feat(KrullDimension): krullDim_eq_iSup_height_add_coheight_of_nonempty (#19155) Part of #15524. This adds a characterization of the Krull Dimenion based on the sum of height and coheight of elements. From the Carleson project. --- Mathlib/Order/KrullDimension.lean | 44 ++++++++++++++++++++++++++----- 1 file changed, 38 insertions(+), 6 deletions(-) diff --git a/Mathlib/Order/KrullDimension.lean b/Mathlib/Order/KrullDimension.lean index 691b858cf4c86..836303c212832 100644 --- a/Mathlib/Order/KrullDimension.lean +++ b/Mathlib/Order/KrullDimension.lean @@ -544,6 +544,15 @@ lemma krullDim_eq_of_orderIso (f : α ≃o β) : krullDim α = krullDim β := le_antisymm (iSup_le fun i ↦ le_sSup ⟨i.reverse, rfl⟩) <| iSup_le fun i ↦ le_sSup ⟨i.reverse, rfl⟩ +lemma height_le_krullDim (a : α) : height a ≤ krullDim α := by + have : Nonempty α := ⟨a⟩ + rw [krullDim_eq_iSup_length] + simp only [WithBot.coe_le_coe, iSup_le_iff] + exact height_le fun p _ ↦ le_iSup_of_le p le_rfl + +lemma coheight_le_krullDim (a : α) : coheight a ≤ krullDim α := by + simpa using height_le_krullDim (α := αᵒᵈ) a + /-- The Krull dimension is the supremum of the elements' heights. @@ -560,10 +569,9 @@ lemma krullDim_eq_iSup_height_of_nonempty [Nonempty α] : krullDim α = ↑(⨆ suffices p.length ≤ ⨆ (a : α), height a by exact (WithBot.unbot'_le_iff fun _ => this).mp this apply le_iSup_of_le p.last (length_le_height_last (p := p)) - · rw [krullDim_eq_iSup_length] - simp only [WithBot.coe_le_coe, iSup_le_iff] - intro x - exact height_le fun p _ ↦ le_iSup_of_le p le_rfl + · rw [WithBot.coe_iSup (by bddDefault)] + apply iSup_le + apply height_le_krullDim /-- The Krull dimension is the supremum of the elements' coheights. @@ -576,8 +584,32 @@ version, with the coercion under the supremum. -/ lemma krullDim_eq_iSup_coheight_of_nonempty [Nonempty α] : krullDim α = ↑(⨆ (a : α), coheight a) := by - rw [← krullDim_orderDual] - exact krullDim_eq_iSup_height_of_nonempty (α := αᵒᵈ) + simpa using krullDim_eq_iSup_height_of_nonempty (α := αᵒᵈ) + +/-- +The Krull dimension is the supremum of the elements' height plus coheight. +-/ +lemma krullDim_eq_iSup_height_add_coheight_of_nonempty [Nonempty α] : + krullDim α = ↑(⨆ (a : α), height a + coheight a) := by + apply le_antisymm + · rw [krullDim_eq_iSup_height_of_nonempty, WithBot.coe_le_coe] + apply ciSup_mono (by bddDefault) (by simp) + · wlog hnottop : krullDim α < ⊤ + · simp_all + rw [krullDim_eq_iSup_length, WithBot.coe_le_coe] + apply iSup_le + intro a + have : height a < ⊤ := WithBot.coe_lt_coe.mp (lt_of_le_of_lt (height_le_krullDim a) hnottop) + have : coheight a < ⊤ := WithBot.coe_lt_coe.mp (lt_of_le_of_lt (coheight_le_krullDim a) hnottop) + cases hh : height a with + | top => simp_all + | coe n => + cases hch : coheight a with + | top => simp_all + | coe m => + obtain ⟨p₁, hlast, hlen₁⟩ := exists_series_of_height_eq_coe a hh + obtain ⟨p₂, hhead, hlen₂⟩ := exists_series_of_coheight_eq_coe a hch + apply le_iSup_of_le ((p₁.smash p₂) (by simp [*])) (by simp [*]) /-- The Krull dimension is the supremum of the elements' heights. From 01fc25a265b2d7c926cefe12201483ab2c41dd82 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 19 Nov 2024 09:31:54 +0000 Subject: [PATCH 104/829] chore(Algebra/Bilinear): deprecate duplicated injectivity lemmas (#19231) These are less general versions of results stated elsewhere. This partially deprecates @Vierkantor's https://github.com/leanprover-community/mathlib3/pull/6588. Also tidy the one caller of these lemmas. --- Mathlib/Algebra/Algebra/Bilinear.lean | 23 ++++++++++------------- Mathlib/RingTheory/Ideal/Basis.lean | 4 ++-- 2 files changed, 12 insertions(+), 15 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Bilinear.lean b/Mathlib/Algebra/Algebra/Bilinear.lean index 0ffbf39951dba..02c54d00368ef 100644 --- a/Mathlib/Algebra/Algebra/Bilinear.lean +++ b/Mathlib/Algebra/Algebra/Bilinear.lean @@ -225,22 +225,19 @@ section Ring variable {R A : Type*} [CommSemiring R] [Ring A] [Algebra R A] +@[deprecated mul_right_injective₀ (since := "2024-11-18")] theorem mulLeft_injective [NoZeroDivisors A] {x : A} (hx : x ≠ 0) : - Function.Injective (mulLeft R x) := by - letI : Nontrivial A := ⟨⟨x, 0, hx⟩⟩ - letI := NoZeroDivisors.to_isDomain A - exact mul_right_injective₀ hx + Function.Injective (mulLeft R x) := + mul_right_injective₀ hx +@[deprecated mul_left_injective₀ (since := "2024-11-18")] theorem mulRight_injective [NoZeroDivisors A] {x : A} (hx : x ≠ 0) : - Function.Injective (mulRight R x) := by - letI : Nontrivial A := ⟨⟨x, 0, hx⟩⟩ - letI := NoZeroDivisors.to_isDomain A - exact mul_left_injective₀ hx - -theorem mul_injective [NoZeroDivisors A] {x : A} (hx : x ≠ 0) : Function.Injective (mul R A x) := by - letI : Nontrivial A := ⟨⟨x, 0, hx⟩⟩ - letI := NoZeroDivisors.to_isDomain A - exact mul_right_injective₀ hx + Function.Injective (mulRight R x) := + mul_left_injective₀ hx + +@[deprecated mul_right_injective₀ (since := "2024-11-18")] +theorem mul_injective [NoZeroDivisors A] {x : A} (hx : x ≠ 0) : Function.Injective (mul R A x) := + mul_right_injective₀ hx end Ring diff --git a/Mathlib/RingTheory/Ideal/Basis.lean b/Mathlib/RingTheory/Ideal/Basis.lean index 4453148c5170a..ac26138204089 100644 --- a/Mathlib/RingTheory/Ideal/Basis.lean +++ b/Mathlib/RingTheory/Ideal/Basis.lean @@ -21,7 +21,7 @@ variable {ι R S : Type*} [CommSemiring R] [CommRing S] [IsDomain S] [Algebra R noncomputable def basisSpanSingleton (b : Basis ι R S) {x : S} (hx : x ≠ 0) : Basis ι R (span ({x} : Set S)) := b.map <| - LinearEquiv.ofInjective (Algebra.lmul R S x) (LinearMap.mul_injective hx) ≪≫ₗ + LinearEquiv.ofInjective (LinearMap.mulLeft R x) (mul_right_injective₀ hx) ≪≫ₗ LinearEquiv.ofEq _ _ (by ext @@ -33,7 +33,7 @@ theorem basisSpanSingleton_apply (b : Basis ι R S) {x : S} (hx : x ≠ 0) (i : (basisSpanSingleton b hx i : S) = x * b i := by simp only [basisSpanSingleton, Basis.map_apply, LinearEquiv.trans_apply, Submodule.restrictScalarsEquiv_apply, LinearEquiv.ofInjective_apply, LinearEquiv.coe_ofEq_apply, - LinearEquiv.restrictScalars_apply, Algebra.coe_lmul_eq_mul, LinearMap.mul_apply'] + LinearEquiv.restrictScalars_apply, LinearMap.mulLeft_apply, LinearMap.mul_apply'] @[simp] theorem constr_basisSpanSingleton {N : Type*} [Semiring N] [Module N S] [SMulCommClass R N S] From 487b33d7af3ccc73184c70deb15c976898d9e5c8 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 19 Nov 2024 10:10:59 +0000 Subject: [PATCH 105/829] chore: fix some left/right injectivity names (#19239) Follow-up Mathlib fixes from [zulip discussion](https://leanprover.zulipchat.com/#narrow/channel/113488-general/topic/.60add_right_inj.60/near/482835891). --- Mathlib/Algebra/Group/Invertible/Defs.lean | 12 ++++++------ Mathlib/Algebra/GroupWithZero/Basic.lean | 4 ++-- Mathlib/Analysis/SpecialFunctions/Bernstein.lean | 4 ++-- 3 files changed, 10 insertions(+), 10 deletions(-) diff --git a/Mathlib/Algebra/Group/Invertible/Defs.lean b/Mathlib/Algebra/Group/Invertible/Defs.lean index 58a3021aca712..44a1421f131e7 100644 --- a/Mathlib/Algebra/Group/Invertible/Defs.lean +++ b/Mathlib/Algebra/Group/Invertible/Defs.lean @@ -230,23 +230,23 @@ section variable [Monoid α] {a b c : α} [Invertible c] variable (c) in -theorem mul_right_inj_of_invertible : a * c = b * c ↔ a = b := +theorem mul_left_inj_of_invertible : a * c = b * c ↔ a = b := ⟨fun h => by simpa using congr_arg (· * ⅟c) h, congr_arg (· * _)⟩ variable (c) in -theorem mul_left_inj_of_invertible : c * a = c * b ↔ a = b := +theorem mul_right_inj_of_invertible : c * a = c * b ↔ a = b := ⟨fun h => by simpa using congr_arg (⅟c * ·) h, congr_arg (_ * ·)⟩ theorem invOf_mul_eq_iff_eq_mul_left : ⅟c * a = b ↔ a = c * b := by - rw [← mul_left_inj_of_invertible (c := c), mul_invOf_cancel_left] + rw [← mul_right_inj_of_invertible (c := c), mul_invOf_cancel_left] theorem mul_left_eq_iff_eq_invOf_mul : c * a = b ↔ a = ⅟c * b := by - rw [← mul_left_inj_of_invertible (c := ⅟c), invOf_mul_cancel_left] + rw [← mul_right_inj_of_invertible (c := ⅟c), invOf_mul_cancel_left] theorem mul_invOf_eq_iff_eq_mul_right : a * ⅟c = b ↔ a = b * c := by - rw [← mul_right_inj_of_invertible (c := c), invOf_mul_cancel_right] + rw [← mul_left_inj_of_invertible (c := c), invOf_mul_cancel_right] theorem mul_right_eq_iff_eq_mul_invOf : a * c = b ↔ a = b * ⅟c := by - rw [← mul_right_inj_of_invertible (c := ⅟c), mul_invOf_cancel_right] + rw [← mul_left_inj_of_invertible (c := ⅟c), mul_invOf_cancel_right] end diff --git a/Mathlib/Algebra/GroupWithZero/Basic.lean b/Mathlib/Algebra/GroupWithZero/Basic.lean index c09c2c9f7742f..f242f9558df45 100644 --- a/Mathlib/Algebra/GroupWithZero/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Basic.lean @@ -244,11 +244,11 @@ section GroupWithZero variable [GroupWithZero G₀] {a b x : G₀} -theorem GroupWithZero.mul_left_injective (h : x ≠ 0) : +theorem GroupWithZero.mul_right_injective (h : x ≠ 0) : Function.Injective fun y => x * y := fun y y' w => by simpa only [← mul_assoc, inv_mul_cancel₀ h, one_mul] using congr_arg (fun y => x⁻¹ * y) w -theorem GroupWithZero.mul_right_injective (h : x ≠ 0) : +theorem GroupWithZero.mul_left_injective (h : x ≠ 0) : Function.Injective fun y => y * x := fun y y' w => by simpa only [mul_assoc, mul_inv_cancel₀ h, mul_one] using congr_arg (fun y => y * x⁻¹) w diff --git a/Mathlib/Analysis/SpecialFunctions/Bernstein.lean b/Mathlib/Analysis/SpecialFunctions/Bernstein.lean index a38081bf15582..4f6e5c73f04c5 100644 --- a/Mathlib/Analysis/SpecialFunctions/Bernstein.lean +++ b/Mathlib/Analysis/SpecialFunctions/Bernstein.lean @@ -109,8 +109,8 @@ theorem probability (n : ℕ) (x : I) : (∑ k : Fin (n + 1), bernstein n k x) = theorem variance {n : ℕ} (h : 0 < (n : ℝ)) (x : I) : (∑ k : Fin (n + 1), (x - k/ₙ : ℝ) ^ 2 * bernstein n k x) = (x : ℝ) * (1 - x) / n := by have h' : (n : ℝ) ≠ 0 := ne_of_gt h - apply_fun fun x : ℝ => x * n using GroupWithZero.mul_right_injective h' - apply_fun fun x : ℝ => x * n using GroupWithZero.mul_right_injective h' + apply_fun fun x : ℝ => x * n using GroupWithZero.mul_left_injective h' + apply_fun fun x : ℝ => x * n using GroupWithZero.mul_left_injective h' dsimp conv_lhs => simp only [Finset.sum_mul, z] conv_rhs => rw [div_mul_cancel₀ _ h'] From 60930c119a77b2c3ae18d4653cb816889fce70a4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 19 Nov 2024 11:04:08 +0000 Subject: [PATCH 106/829] chore: shortcut instances for `Finite Bool`, `Finite Prop` (#19086) These can be defined much earlier. Also add `WellFoundedLT` shortcut instances. [Zulip](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/WellFoundedLT.20Prop.20is.20not.20found.20when.20importing.20too.20much) --- Mathlib/Data/Finite/Defs.lean | 3 +++ Mathlib/Data/Fintype/Card.lean | 7 +++++++ 2 files changed, 10 insertions(+) diff --git a/Mathlib/Data/Finite/Defs.lean b/Mathlib/Data/Finite/Defs.lean index 784e71e12e4f6..dce5c020d7baf 100644 --- a/Mathlib/Data/Finite/Defs.lean +++ b/Mathlib/Data/Finite/Defs.lean @@ -162,6 +162,9 @@ protected theorem Infinite.false [Finite α] (_ : Infinite α) : False := alias ⟨Finite.of_not_infinite, Finite.not_infinite⟩ := not_infinite_iff_finite +instance Bool.instFinite : Finite Bool := .intro finTwoEquiv.symm +instance Prop.instFinite : Finite Prop := .of_equiv _ Equiv.propEquivBool.symm + section Set /-! diff --git a/Mathlib/Data/Fintype/Card.lean b/Mathlib/Data/Fintype/Card.lean index dc1e60e25f55d..59e57330eb779 100644 --- a/Mathlib/Data/Fintype/Card.lean +++ b/Mathlib/Data/Fintype/Card.lean @@ -841,6 +841,13 @@ instance (priority := 100) to_wellFoundedGT [Preorder α] : WellFoundedGT α := end Finite +-- Shortcut instances to make sure those are found even in the presence of other instances +-- See https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/WellFoundedLT.20Prop.20is.20not.20found.20when.20importing.20too.20much +instance Bool.instWellFoundedLT : WellFoundedLT Bool := inferInstance +instance Bool.instWellFoundedGT : WellFoundedGT Bool := inferInstance +instance Prop.instWellFoundedLT : WellFoundedLT Prop := inferInstance +instance Prop.instWellFoundedGT : WellFoundedGT Prop := inferInstance + -- @[nolint fintype_finite] -- Porting note: do we need this? protected theorem Fintype.false [Infinite α] (_h : Fintype α) : False := not_finite α From 2fd09beb18e05ada8bff95d82fdc2e0e6a55ba0a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Tom=C3=A1=C5=A1=20Sk=C5=99ivan?= Date: Tue, 19 Nov 2024 11:04:10 +0000 Subject: [PATCH 107/829] feat(fun_prop): failure cache for `fun_prop` and command to print out `fun_prop` theorems (#19112) Add new cache to `fun_prop` for failed goals. `fun_prop` used to take exponentially long to fail without this cache. Also add command `#print_fun_prop_theorems HAdd.hAdd` prints out all `fun_prop` theorems associated with addition. --- Mathlib/Tactic/FunProp/Core.lean | 50 +++++++++++------ Mathlib/Tactic/FunProp/Elab.lean | 53 ++++++++++++++++++ Mathlib/Tactic/FunProp/Theorems.lean | 2 +- Mathlib/Tactic/FunProp/Types.lean | 4 +- MathlibTest/fun_prop_dev.lean | 83 +++++++++++++++++++++++++--- 5 files changed, 166 insertions(+), 26 deletions(-) diff --git a/Mathlib/Tactic/FunProp/Core.lean b/Mathlib/Tactic/FunProp/Core.lean index 7cfc38fde4b73..9f84b5b2c1371 100644 --- a/Mathlib/Tactic/FunProp/Core.lean +++ b/Mathlib/Tactic/FunProp/Core.lean @@ -507,20 +507,18 @@ def tryTheorems (funPropDecl : FunPropDecl) (e : Expr) (fData : FunctionData) if let .some r ← tryTheorem? e thm.thmOrigin funProp then return r | .some (.some (f,g)) => - trace[Debug.Meta.Tactic.fun_prop] - s!"decomposing to later use {←ppOrigin' thm.thmOrigin}" - trace[Debug.Meta.Tactic.fun_prop] - s!"decomposition: {← ppExpr f} ∘ {← ppExpr g}" + trace[Meta.Tactic.fun_prop] + s!"decomposing to later use {←ppOrigin' thm.thmOrigin} as: + ({← ppExpr f}) ∘ ({← ppExpr g})" if let .some r ← applyCompRule funPropDecl e f g funProp then return r | _ => continue else - trace[Debug.Meta.Tactic.fun_prop] - s!"decomposing in args {thm.mainArgs} to later use {←ppOrigin' thm.thmOrigin}" let .some (f,g) ← fData.decompositionOverArgs thm.mainArgs | continue - trace[Debug.Meta.Tactic.fun_prop] - s!"decomposition: {← ppExpr f} ∘ {← ppExpr g}" + trace[Meta.Tactic.fun_prop] + s!"decomposing to later use {←ppOrigin' thm.thmOrigin} as: + ({← ppExpr f}) ∘ ({← ppExpr g})" if let .some r ← applyCompRule funPropDecl e f g funProp then return r -- todo: decompose if uncurried and arguments do not match exactly @@ -594,9 +592,17 @@ def constAppCase (funPropDecl : FunPropDecl) (e : Expr) (fData : FunctionData) return r if let .some (f,g) ← fData.nontrivialDecomposition then + trace[Meta.Tactic.fun_prop] + s!"failed applying `{funPropDecl.funPropName}` theorems for `{funName}` + trying again after decomposing function as: `({← ppExpr f}) ∘ ({← ppExpr g})`" + if let .some r ← applyCompRule funPropDecl e f g funProp then return r else + trace[Meta.Tactic.fun_prop] + s!"failed applying `{funPropDecl.funPropName}` theorems for `{funName}` + now trying to prove `{funPropDecl.funPropName}` from another function property" + if let .some r ← applyTransitionRules e funProp then return r @@ -609,15 +615,27 @@ def cacheResult (e : Expr) (r : Result) : FunPropM Result := do -- return proof? modify (fun s => { s with cache := s.cache.insert e { expr := q(True), proof? := r.proof} }) return r +/-- Cache for failed goals such that `fun_prop` can fail fast next time. -/ +def cacheFailure (e : Expr) : FunPropM Unit := do -- return proof? + modify (fun s => { s with failureCache := s.failureCache.insert e }) + + mutual /-- Main `funProp` function. Returns proof of `e`. -/ partial def funProp (e : Expr) : FunPropM (Option Result) := do let e ← instantiateMVars e - -- check cache + + withTraceNode `Meta.Tactic.fun_prop + (fun r => do pure s!"[{ExceptToEmoji.toEmoji r}] {← ppExpr e}") do + + -- check cache for succesfull goals if let .some { expr := _, proof? := .some proof } := (← get).cache.find? e then - trace[Debug.Meta.Tactic.fun_prop] "cached result for {e}" + trace[Meta.Tactic.fun_prop] "reusing previously found proof for {e}" return .some { proof := proof } + else if (← get).failureCache.contains e then + trace[Meta.Tactic.fun_prop] "skipping proof search, proving {e} was tried already and failed" + return .none else -- take care of forall and let binders and run main match e with @@ -633,9 +651,12 @@ mutual cacheResult e {proof := ← mkLambdaFVars xs r.proof } | .mdata _ e' => funProp e' | _ => - let .some r ← main e - | return none - cacheResult e r + if let .some r ← main e then + cacheResult e r + else + cacheFailure e + return none + /-- Main `funProp` function. Returns proof of `e`. -/ private partial def main (e : Expr) : FunPropM (Option Result) := do @@ -645,9 +666,6 @@ mutual increaseSteps - withTraceNode `Meta.Tactic.fun_prop - (fun r => do pure s!"[{ExceptToEmoji.toEmoji r}] {← ppExpr e}") do - -- if function starts with let bindings move them the top of `e` and try again if f.isLet then return ← letTelescope f fun xs b => do diff --git a/Mathlib/Tactic/FunProp/Elab.lean b/Mathlib/Tactic/FunProp/Elab.lean index 94afec84e9177..6761d53f837c6 100644 --- a/Mathlib/Tactic/FunProp/Elab.lean +++ b/Mathlib/Tactic/FunProp/Elab.lean @@ -83,6 +83,59 @@ def funPropTac : Tactic | _ => throwUnsupportedSyntax + + +/-- Command that printins all function properties attached to a function. + +For example +``` +#print_fun_prop_theorems HAdd.hAdd +``` +might print out +``` +Continuous + continuous_add, args: [4,5], priority: 1000 + continuous_add_left, args: [5], priority: 1000 + continuous_add_right, args [4], priority: 1000 + ... +Diferentiable + Differentiable.add, args: [4,5], priority: 1000 + Differentiable.add_const, args: [4], priority: 1000 + Differentiable.const_add, args: [5], priority: 1000 + ... +``` + +You can also see only theorems about a concrete function property +``` +#print_fun_prop_theorems HAdd.hAdd Continuous +``` +-/ +elab "#print_fun_prop_theorems " funIdent:ident funProp:(ident)? : command => do + + let funName ← ensureNonAmbiguous funIdent (← resolveGlobalConst funIdent) + let funProp? ← funProp.mapM (fun stx => do + ensureNonAmbiguous stx (← resolveGlobalConst stx)) + + let theorems := (functionTheoremsExt.getState (← getEnv)).theorems.findD funName {} + + let logTheorems (funProp : Name) (thms : Array FunctionTheorem) : Command.CommandElabM Unit := do + let mut msg : MessageData := "" + msg := msg ++ m!"{← Meta.ppOrigin (.decl funProp)}" + for thm in thms do + msg := msg ++ m!"\n {← Meta.ppOrigin (.decl thm.thmOrigin.name)}, \ + args: {thm.mainArgs}, form: {thm.form}" + pure () + logInfo msg + + + match funProp? with + | none => + for (funProp,thms) in theorems do + logTheorems funProp thms + | some funProp => + logTheorems funProp (theorems.findD funProp #[]) + + end Meta.FunProp end Mathlib diff --git a/Mathlib/Tactic/FunProp/Theorems.lean b/Mathlib/Tactic/FunProp/Theorems.lean index 1ab9b3850fec5..b0c184aedc988 100644 --- a/Mathlib/Tactic/FunProp/Theorems.lean +++ b/Mathlib/Tactic/FunProp/Theorems.lean @@ -145,7 +145,7 @@ inductive TheoremForm where /-- TheoremForm to string -/ instance : ToString TheoremForm := - ⟨fun x => match x with | .uncurried => "uncurried" | .comp => "compositional"⟩ + ⟨fun x => match x with | .uncurried => "simple" | .comp => "compositional"⟩ /-- theorem about specific function (either declared constant or free variable) -/ structure FunctionTheorem where diff --git a/Mathlib/Tactic/FunProp/Types.lean b/Mathlib/Tactic/FunProp/Types.lean index 6f12b788e1928..a8af67ac13b91 100644 --- a/Mathlib/Tactic/FunProp/Types.lean +++ b/Mathlib/Tactic/FunProp/Types.lean @@ -97,8 +97,10 @@ structure Context where /-- `fun_prop` state -/ structure State where /-- Simp's cache is used as the `fun_prop` tactic is designed to be used inside of simp and - utilize its cache -/ + utilize its cache. It holds successful goals. -/ cache : Simp.Cache := {} + /-- Cache storing failed goals such that they are not tried again. -/ + failureCache : ExprSet := {} /-- Count the number of steps and stop when maxSteps is reached. -/ numSteps := 0 /-- Log progress and failures messages that should be displayed to the user at the end. -/ diff --git a/MathlibTest/fun_prop_dev.lean b/MathlibTest/fun_prop_dev.lean index f78abae43617a..866e3b6aebefa 100644 --- a/MathlibTest/fun_prop_dev.lean +++ b/MathlibTest/fun_prop_dev.lean @@ -6,6 +6,7 @@ Authors: Tomáš Skřivan import Mathlib.Tactic.FunProp import Mathlib.Logic.Function.Basic import Mathlib.Data.FunLike.Basic +import Mathlib.Tactic.SuccessIfFailWithMsg import Aesop /-! # Tests for the `fun_prop` tactic @@ -536,14 +537,15 @@ example [Add α] (a : α) : -- Test that local theorem is being used /-- info: [Meta.Tactic.fun_prop] [✅️] Con fun x => f x y - [Meta.Tactic.fun_prop] candidate local theorems for f #[this : Con f] - [Meta.Tactic.fun_prop] removing argument to later use this : Con f - [Meta.Tactic.fun_prop] [✅️] applying: Con_comp - [Meta.Tactic.fun_prop] [✅️] Con fun f => f y - [Meta.Tactic.fun_prop] [✅️] applying: Con_apply - [Meta.Tactic.fun_prop] [✅️] Con fun x => f x - [Meta.Tactic.fun_prop] candidate local theorems for f #[this : Con f] - [Meta.Tactic.fun_prop] [✅️] applying: this : Con f + [Meta.Tactic.fun_prop] [✅️] Con fun x => f x y + [Meta.Tactic.fun_prop] candidate local theorems for f #[this : Con f] + [Meta.Tactic.fun_prop] removing argument to later use this : Con f + [Meta.Tactic.fun_prop] [✅️] applying: Con_comp + [Meta.Tactic.fun_prop] [✅️] Con fun f => f y + [Meta.Tactic.fun_prop] [✅️] applying: Con_apply + [Meta.Tactic.fun_prop] [✅️] Con fun x => f x + [Meta.Tactic.fun_prop] candidate local theorems for f #[this : Con f] + [Meta.Tactic.fun_prop] [✅️] applying: this : Con f -/ #guard_msgs in example [Add α] (y : α): @@ -553,3 +555,68 @@ example [Add α] (y : α): have : Con f := by fun_prop set_option trace.Meta.Tactic.fun_prop true in fun_prop + + + +--- pefromance tests - mainly testing fast failure --- +------------------------------------------------------ + + +section PerformanceTests +-- set_option trace.Meta.Tactic.fun_prop true +-- set_option profiler true + +variable {R : Type*} [Add R] [∀ n, OfNat R n] +example (f : R → R) (hf : Con f) : + Con (fun x ↦ f (x + 3)) := by fun_prop -- succeeds in 5ms +example (f : R → R) (hf : Con f) : + Con (fun x ↦ (f (x + 3)) + 2) := by fun_prop -- succeeds in 6ms +example (f : R → R) (hf : Con f) : + Con (fun x ↦ (f (x + 3)) + (2 + f (x + 1))) := by fun_prop -- succeeds in 8ms +example (f : R → R) (hf : Con f) : + Con (fun x ↦ (f (x + 3)) + (2 + f (x + 1)) + x) := by fun_prop -- succeeds in 10ms +example (f : R → R) (hf : Con f) : + Con (fun x ↦ (f (x + 3)) + 2 + f (x + 1) + x + 1) := by fun_prop -- succeeds in 11ms + +-- This used to fail in exponentially increasing time, up to 6s for the last example +-- We set maxHearthbeats to 1000 such that the last three examples should fail if the exponential +-- blow happen again. +set_option maxHeartbeats 1000 in +example (f : R → R) : + Con (fun x ↦ f (x + 3)) := by + fail_if_success fun_prop + apply silentSorry + +example (f : R → R) : + Con (fun x ↦ (f (x + 3)) + 2) := by + fail_if_success fun_prop + apply silentSorry + +set_option maxHeartbeats 1000 in +example (f : R → R) : + Con (fun x ↦ ((f (x + 3)) + 2) + f (x + 1)) := by + fail_if_success fun_prop + apply silentSorry + +set_option maxHeartbeats 1000 in +example (f : R → R) : + Con (fun x ↦ ((f (x + 3)) + 2) + f (x + 1) + x) := by + fail_if_success fun_prop + apply silentSorry + +set_option maxHeartbeats 1000 in +example (f : R → R) : + Con (fun x ↦ ((f (x + 3)) + 2) + f (x + 1) + x + 1) := by + fail_if_success fun_prop + apply silentSorry + +end PerformanceTests + + +/-- +info: Con + add_Con, args: [4, 5], form: simple + add_Con', args: [4, 5], form: compositional +-/ +#guard_msgs in +#print_fun_prop_theorems HAdd.hAdd Con From 1212c0adb03d7db70862a182885d0cf7ddd9dfed Mon Sep 17 00:00:00 2001 From: grunweg Date: Tue, 19 Nov 2024 11:04:11 +0000 Subject: [PATCH 108/829] chore: some renames `iff_open` -> `iff_isOpen` (#19209) Our naming convention prefers `isOpen`. See the individual commit messages for details. --- .../Constructions/BorelSpace/Basic.lean | 2 +- .../Topology/Algebra/UniformGroup/Defs.lean | 7 ++++++- .../Topology/Connected/LocallyConnected.lean | 18 ++++++++++++------ Mathlib/Topology/Defs/Filter.lean | 4 ++-- Mathlib/Topology/Inseparable.lean | 15 ++++++++++----- Mathlib/Topology/Separation/Basic.lean | 4 ++-- Mathlib/Topology/Sets/Opens.lean | 2 +- Mathlib/Topology/UniformSpace/Basic.lean | 7 +++++-- 8 files changed, 39 insertions(+), 20 deletions(-) diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean index de06be836259b..23adc61b8b596 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean @@ -330,7 +330,7 @@ instance (priority := 100) OpensMeasurableSpace.separatesPoints [T0Space α] : rw [separatesPoints_iff] intro x y hxy apply Inseparable.eq - rw [inseparable_iff_forall_open] + rw [inseparable_iff_forall_isOpen] exact fun s hs => hxy _ hs.measurableSet theorem borel_eq_top_of_countable {α : Type*} [TopologicalSpace α] [T0Space α] [Countable α] : diff --git a/Mathlib/Topology/Algebra/UniformGroup/Defs.lean b/Mathlib/Topology/Algebra/UniformGroup/Defs.lean index 1ebc4e5aa7b0d..b593b9d1eb3ca 100644 --- a/Mathlib/Topology/Algebra/UniformGroup/Defs.lean +++ b/Mathlib/Topology/Algebra/UniformGroup/Defs.lean @@ -324,7 +324,7 @@ theorem MonoidHom.uniformContinuous_of_continuousAt_one [UniformSpace β] [Group its kernel is open. -/ @[to_additive "A homomorphism from a uniform additive group to a discrete uniform additive group is continuous if and only if its kernel is open."] -theorem UniformGroup.uniformContinuous_iff_open_ker {hom : Type*} [UniformSpace β] +theorem UniformGroup.uniformContinuous_iff_isOpen_ker {hom : Type*} [UniformSpace β] [DiscreteTopology β] [Group β] [UniformGroup β] [FunLike hom α β] [MonoidHomClass hom α β] {f : hom} : UniformContinuous f ↔ IsOpen ((f : α →* β).ker : Set α) := by @@ -334,6 +334,11 @@ theorem UniformGroup.uniformContinuous_iff_open_ker {hom : Type*} [UniformSpace rw [ContinuousAt, nhds_discrete β, map_one, tendsto_pure] exact hf.mem_nhds (map_one f) +@[deprecated (since := "2024-11-18")] alias UniformGroup.uniformContinuous_iff_open_ker := + UniformGroup.uniformContinuous_iff_isOpen_ker +@[deprecated (since := "2024-11-18")] alias UniformAddGroup.uniformContinuous_iff_open_ker := + UniformAddGroup.uniformContinuous_iff_isOpen_ker + @[to_additive] theorem uniformContinuous_monoidHom_of_continuous {hom : Type*} [UniformSpace β] [Group β] [UniformGroup β] [FunLike hom α β] [MonoidHomClass hom α β] {f : hom} (h : Continuous f) : diff --git a/Mathlib/Topology/Connected/LocallyConnected.lean b/Mathlib/Topology/Connected/LocallyConnected.lean index ff3a0ebf66618..2e690ee0fbb2d 100644 --- a/Mathlib/Topology/Connected/LocallyConnected.lean +++ b/Mathlib/Topology/Connected/LocallyConnected.lean @@ -31,15 +31,18 @@ class LocallyConnectedSpace (α : Type*) [TopologicalSpace α] : Prop where /-- Open connected neighborhoods form a basis of the neighborhoods filter. -/ open_connected_basis : ∀ x, (𝓝 x).HasBasis (fun s : Set α => IsOpen s ∧ x ∈ s ∧ IsConnected s) id -theorem locallyConnectedSpace_iff_open_connected_basis : +theorem locallyConnectedSpace_iff_hasBasis_isOpen_isConnected : LocallyConnectedSpace α ↔ ∀ x, (𝓝 x).HasBasis (fun s : Set α => IsOpen s ∧ x ∈ s ∧ IsConnected s) id := ⟨@LocallyConnectedSpace.open_connected_basis _ _, LocallyConnectedSpace.mk⟩ -theorem locallyConnectedSpace_iff_open_connected_subsets : +@[deprecated (since := "2024-11-18")] alias locallyConnectedSpace_iff_open_connected_basis := + locallyConnectedSpace_iff_hasBasis_isOpen_isConnected + +theorem locallyConnectedSpace_iff_subsets_isOpen_isConnected : LocallyConnectedSpace α ↔ ∀ x, ∀ U ∈ 𝓝 x, ∃ V : Set α, V ⊆ U ∧ IsOpen V ∧ x ∈ V ∧ IsConnected V := by - simp_rw [locallyConnectedSpace_iff_open_connected_basis] + simp_rw [locallyConnectedSpace_iff_hasBasis_isOpen_isConnected] refine forall_congr' fun _ => ?_ constructor · intro h U hU @@ -49,10 +52,13 @@ theorem locallyConnectedSpace_iff_open_connected_subsets : let ⟨V, hVU, hV⟩ := h U hU ⟨V, hV, hVU⟩, fun ⟨V, ⟨hV, hxV, _⟩, hVU⟩ => mem_nhds_iff.mpr ⟨V, hVU, hV, hxV⟩⟩⟩ +@[deprecated (since := "2024-11-18")] alias locallyConnectedSpace_iff_open_connected_subsets := + locallyConnectedSpace_iff_subsets_isOpen_isConnected + /-- A space with discrete topology is a locally connected space. -/ instance (priority := 100) DiscreteTopology.toLocallyConnectedSpace (α) [TopologicalSpace α] [DiscreteTopology α] : LocallyConnectedSpace α := - locallyConnectedSpace_iff_open_connected_subsets.2 fun x _U hU => + locallyConnectedSpace_iff_subsets_isOpen_isConnected.2 fun x _U hU => ⟨{x}, singleton_subset_iff.2 <| mem_of_mem_nhds hU, isOpen_discrete _, rfl, isConnected_singleton⟩ @@ -85,7 +91,7 @@ theorem locallyConnectedSpace_iff_connectedComponentIn_open : · intro h exact fun F hF x _ => hF.connectedComponentIn · intro h - rw [locallyConnectedSpace_iff_open_connected_subsets] + rw [locallyConnectedSpace_iff_subsets_isOpen_isConnected] refine fun x U hU => ⟨connectedComponentIn (interior U) x, (connectedComponentIn_subset _ _).trans interior_subset, h _ isOpen_interior x ?_, @@ -95,7 +101,7 @@ theorem locallyConnectedSpace_iff_connectedComponentIn_open : theorem locallyConnectedSpace_iff_connected_subsets : LocallyConnectedSpace α ↔ ∀ (x : α), ∀ U ∈ 𝓝 x, ∃ V ∈ 𝓝 x, IsPreconnected V ∧ V ⊆ U := by constructor - · rw [locallyConnectedSpace_iff_open_connected_subsets] + · rw [locallyConnectedSpace_iff_subsets_isOpen_isConnected] intro h x U hxU rcases h x U hxU with ⟨V, hVU, hV₁, hxV, hV₂⟩ exact ⟨V, hV₁.mem_nhds hxV, hV₂.isPreconnected, hVU⟩ diff --git a/Mathlib/Topology/Defs/Filter.lean b/Mathlib/Topology/Defs/Filter.lean index 462aaeabab057..340ed360bb1da 100644 --- a/Mathlib/Topology/Defs/Filter.lean +++ b/Mathlib/Topology/Defs/Filter.lean @@ -202,8 +202,8 @@ infixl:300 " ⤳ " => Specializes equivalent properties hold: - `𝓝 x = 𝓝 y`; we use this property as the definition; -- for any open set `s`, `x ∈ s ↔ y ∈ s`, see `inseparable_iff_open`; -- for any closed set `s`, `x ∈ s ↔ y ∈ s`, see `inseparable_iff_closed`; +- for any open set `s`, `x ∈ s ↔ y ∈ s`, see `inseparable_iff_forall_isOpen`; +- for any closed set `s`, `x ∈ s ↔ y ∈ s`, see `inseparable_iff_forall_isClosed`; - `x ∈ closure {y}` and `y ∈ closure {x}`, see `inseparable_iff_mem_closure`; - `closure {x} = closure {y}`, see `inseparable_iff_closure_eq`. -/ diff --git a/Mathlib/Topology/Inseparable.lean b/Mathlib/Topology/Inseparable.lean index 2c7156a1d00f3..7b0bf5a54bf81 100644 --- a/Mathlib/Topology/Inseparable.lean +++ b/Mathlib/Topology/Inseparable.lean @@ -433,17 +433,22 @@ theorem Inseparable.specializes' (h : x ~ᵢ y) : y ⤳ x := h.ge theorem Specializes.antisymm (h₁ : x ⤳ y) (h₂ : y ⤳ x) : x ~ᵢ y := le_antisymm h₁ h₂ -theorem inseparable_iff_forall_open : (x ~ᵢ y) ↔ ∀ s : Set X, IsOpen s → (x ∈ s ↔ y ∈ s) := by +theorem inseparable_iff_forall_isOpen : (x ~ᵢ y) ↔ ∀ s : Set X, IsOpen s → (x ∈ s ↔ y ∈ s) := by simp only [inseparable_iff_specializes_and, specializes_iff_forall_open, ← forall_and, ← iff_def, Iff.comm] +@[deprecated (since := "2024-11-18")] alias +inseparable_iff_forall_open := inseparable_iff_forall_isOpen + theorem not_inseparable_iff_exists_open : ¬(x ~ᵢ y) ↔ ∃ s : Set X, IsOpen s ∧ Xor' (x ∈ s) (y ∈ s) := by - simp [inseparable_iff_forall_open, ← xor_iff_not_iff] + simp [inseparable_iff_forall_isOpen, ← xor_iff_not_iff] -theorem inseparable_iff_forall_closed : (x ~ᵢ y) ↔ ∀ s : Set X, IsClosed s → (x ∈ s ↔ y ∈ s) := by +theorem inseparable_iff_forall_isClosed : (x ~ᵢ y) ↔ ∀ s : Set X, IsClosed s → (x ∈ s ↔ y ∈ s) := by simp only [inseparable_iff_specializes_and, specializes_iff_forall_closed, ← forall_and, ← iff_def] +@[deprecated (since := "2024-11-18")] alias +inseparable_iff_forall_closed := inseparable_iff_forall_isClosed theorem inseparable_iff_mem_closure : (x ~ᵢ y) ↔ x ∈ closure ({y} : Set X) ∧ y ∈ closure ({x} : Set X) := @@ -497,10 +502,10 @@ nonrec theorem trans (h₁ : x ~ᵢ y) (h₂ : y ~ᵢ z) : x ~ᵢ z := h₁.tran theorem nhds_eq (h : x ~ᵢ y) : 𝓝 x = 𝓝 y := h theorem mem_open_iff (h : x ~ᵢ y) (hs : IsOpen s) : x ∈ s ↔ y ∈ s := - inseparable_iff_forall_open.1 h s hs + inseparable_iff_forall_isOpen.1 h s hs theorem mem_closed_iff (h : x ~ᵢ y) (hs : IsClosed s) : x ∈ s ↔ y ∈ s := - inseparable_iff_forall_closed.1 h s hs + inseparable_iff_forall_isClosed.1 h s hs theorem map_of_continuousAt (h : x ~ᵢ y) (hx : ContinuousAt f x) (hy : ContinuousAt f y) : f x ~ᵢ f y := diff --git a/Mathlib/Topology/Separation/Basic.lean b/Mathlib/Topology/Separation/Basic.lean index 7698efb12ea4d..b56a3a27aa9a9 100644 --- a/Mathlib/Topology/Separation/Basic.lean +++ b/Mathlib/Topology/Separation/Basic.lean @@ -316,7 +316,7 @@ theorem inseparable_eq_eq [T0Space X] : Inseparable = @Eq X := theorem TopologicalSpace.IsTopologicalBasis.inseparable_iff {b : Set (Set X)} (hb : IsTopologicalBasis b) {x y : X} : Inseparable x y ↔ ∀ s ∈ b, (x ∈ s ↔ y ∈ s) := - ⟨fun h _ hs ↦ inseparable_iff_forall_open.1 h _ (hb.isOpen hs), + ⟨fun h _ hs ↦ inseparable_iff_forall_isOpen.1 h _ (hb.isOpen hs), fun h ↦ hb.nhds_hasBasis.eq_of_same_basis <| by convert hb.nhds_hasBasis using 2 exact and_congr_right (h _)⟩ @@ -328,7 +328,7 @@ theorem TopologicalSpace.IsTopologicalBasis.eq_iff [T0Space X] {b : Set (Set X)} theorem t0Space_iff_exists_isOpen_xor'_mem (X : Type u) [TopologicalSpace X] : T0Space X ↔ Pairwise fun x y => ∃ U : Set X, IsOpen U ∧ Xor' (x ∈ U) (y ∈ U) := by simp only [t0Space_iff_not_inseparable, xor_iff_not_iff, not_forall, exists_prop, - inseparable_iff_forall_open, Pairwise] + inseparable_iff_forall_isOpen, Pairwise] theorem exists_isOpen_xor'_mem [T0Space X] {x y : X} (h : x ≠ y) : ∃ U : Set X, IsOpen U ∧ Xor' (x ∈ U) (y ∈ U) := diff --git a/Mathlib/Topology/Sets/Opens.lean b/Mathlib/Topology/Sets/Opens.lean index 574a884fa980a..452c6d12ef50e 100644 --- a/Mathlib/Topology/Sets/Opens.lean +++ b/Mathlib/Topology/Sets/Opens.lean @@ -378,7 +378,7 @@ theorem comap_injective [T0Space β] : Injective (comap : C(α, β) → FrameHom fun f g h => ContinuousMap.ext fun a => Inseparable.eq <| - inseparable_iff_forall_open.2 fun s hs => + inseparable_iff_forall_isOpen.2 fun s hs => have : comap f ⟨s, hs⟩ = comap g ⟨s, hs⟩ := DFunLike.congr_fun h ⟨_, hs⟩ show a ∈ f ⁻¹' s ↔ a ∈ g ⁻¹' s from Set.ext_iff.1 (coe_inj.2 this) a diff --git a/Mathlib/Topology/UniformSpace/Basic.lean b/Mathlib/Topology/UniformSpace/Basic.lean index 9bfe9b67a53a3..47f236596aa6b 100644 --- a/Mathlib/Topology/UniformSpace/Basic.lean +++ b/Mathlib/Topology/UniformSpace/Basic.lean @@ -651,7 +651,7 @@ theorem nhdsWithin_eq_comap_uniformity {x : α} (S : Set α) : 𝓝[S] x = (𝓤 α ⊓ 𝓟 (univ ×ˢ S)).comap (Prod.mk x) := nhdsWithin_eq_comap_uniformity_of_mem (mem_univ _) S -/-- See also `isOpen_iff_open_ball_subset`. -/ +/-- See also `isOpen_iff_isOpen_ball_subset`. -/ theorem isOpen_iff_ball_subset {s : Set α} : IsOpen s ↔ ∀ x ∈ s, ∃ V ∈ 𝓤 α, ball x V ⊆ s := by simp_rw [isOpen_iff_mem_nhds, nhds_eq_comap_uniformity, mem_comap, ball] @@ -863,7 +863,7 @@ theorem mem_uniformity_isClosed {s : Set (α × α)} (h : s ∈ 𝓤 α) : ∃ t let ⟨t, ⟨ht_mem, htc⟩, hts⟩ := uniformity_hasBasis_closed.mem_iff.1 h ⟨t, ht_mem, htc, hts⟩ -theorem isOpen_iff_open_ball_subset {s : Set α} : +theorem isOpen_iff_isOpen_ball_subset {s : Set α} : IsOpen s ↔ ∀ x ∈ s, ∃ V ∈ 𝓤 α, IsOpen V ∧ ball x V ⊆ s := by rw [isOpen_iff_ball_subset] constructor <;> intro h x hx @@ -874,6 +874,9 @@ theorem isOpen_iff_open_ball_subset {s : Set α} : · obtain ⟨V, hV, -, hV'⟩ := h x hx exact ⟨V, hV, hV'⟩ +@[deprecated (since := "2024-11-18")] alias +isOpen_iff_open_ball_subset := isOpen_iff_isOpen_ball_subset + /-- The uniform neighborhoods of all points of a dense set cover the whole space. -/ theorem Dense.biUnion_uniformity_ball {s : Set α} {U : Set (α × α)} (hs : Dense s) (hU : U ∈ 𝓤 α) : ⋃ x ∈ s, ball x U = univ := by From 442a60cbe707e04d02c26db93db0f22e7e19391c Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Tue, 19 Nov 2024 11:04:13 +0000 Subject: [PATCH 109/829] chore: update Mathlib dependencies 2024-11-19 (#19242) 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 2464526042495..cfc65dc4caaa9 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "01f4969b6e861db6a99261ea5eadd5a9bb63011b", + "rev": "485efbc439ee0ebdeae8afb0acd24a5e82e2f771", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From f7b68aa4cf35ead3a763da191bba0ff07de3f96c Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Tue, 19 Nov 2024 11:44:13 +0000 Subject: [PATCH 110/829] feat(AlgebraicGeometry): ring hom is integral if the induced map between affine lines is closed (#18804) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib/Algebra/Polynomial/Reverse.lean | 8 ++ .../PrimeSpectrum/Basic.lean | 75 +++++++++++++++++-- 2 files changed, 76 insertions(+), 7 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/Reverse.lean b/Mathlib/Algebra/Polynomial/Reverse.lean index 505b6a0ea5d5a..07d080133f5ee 100644 --- a/Mathlib/Algebra/Polynomial/Reverse.lean +++ b/Mathlib/Algebra/Polynomial/Reverse.lean @@ -143,6 +143,14 @@ theorem reflect_monomial (N n : ℕ) : reflect N ((X : R[X]) ^ n) = X ^ revAt N @[simp] lemma reflect_one_X : reflect 1 (X : R[X]) = 1 := by simpa using reflect_monomial 1 1 (R := R) +lemma reflect_map {S : Type*} [Semiring S] (f : R →+* S) (p : R[X]) (n : ℕ) : + (p.map f).reflect n = (p.reflect n).map f := by + ext; simp + +@[simp] +lemma reflect_one (n : ℕ) : (1 : R[X]).reflect n = Polynomial.X ^ n := by + rw [← C.map_one, reflect_C, map_one, one_mul] + theorem reflect_mul_induction (cf cg : ℕ) : ∀ N O : ℕ, ∀ f g : R[X], diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index df037ab686364..b1fe379b4c1cc 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -321,13 +321,6 @@ theorem comap_singleton_isClosed_of_surjective (f : R →+* S) (hf : Function.Su haveI : x.asIdeal.IsMaximal := (isClosed_singleton_iff_isMaximal x).1 hx (isClosed_singleton_iff_isMaximal _).2 (Ideal.comap_isMaximal_of_surjective f hf) -theorem comap_singleton_isClosed_of_isIntegral (f : R →+* S) (hf : f.IsIntegral) - (x : PrimeSpectrum S) (hx : IsClosed ({x} : Set (PrimeSpectrum S))) : - IsClosed ({comap f x} : Set (PrimeSpectrum R)) := - have := (isClosed_singleton_iff_isMaximal x).1 hx - (isClosed_singleton_iff_isMaximal _).2 - (Ideal.isMaximal_comap_of_isIntegral_of_isMaximal' f hf x.asIdeal) - theorem image_comap_zeroLocus_eq_zeroLocus_comap (hf : Surjective f) (I : Ideal S) : comap f '' zeroLocus I = zeroLocus (I.comap f) := image_specComap_zeroLocus_eq_zeroLocus_comap _ f hf I @@ -719,6 +712,74 @@ lemma exists_idempotent_basicOpen_eq_of_is_clopen {s : Set (PrimeSpectrum R)} exact PrimeSpectrum.zeroLocus_anti_mono (Set.singleton_subset_iff.mpr <| Ideal.pow_le_self hnz hx) +section IsIntegral + +open Polynomial + +variable {R S : Type*} [CommRing R] [CommRing S] (f : R →+* S) + +theorem isClosedMap_comap_of_isIntegral (hf : f.IsIntegral) : + IsClosedMap (comap f) := by + refine fun s hs ↦ isClosed_image_of_stableUnderSpecialization _ _ hs ?_ + rintro _ y e ⟨x, hx, rfl⟩ + algebraize [f] + obtain ⟨q, hq₁, hq₂, hq₃⟩ := Ideal.exists_ideal_over_prime_of_isIntegral y.asIdeal x.asIdeal + ((le_iff_specializes _ _).mpr e) + refine ⟨⟨q, hq₂⟩, ((le_iff_specializes _ ⟨q, hq₂⟩).mp hq₁).mem_closed hs hx, + PrimeSpectrum.ext hq₃⟩ + +theorem isClosed_comap_singleton_of_isIntegral (hf : f.IsIntegral) + (x : PrimeSpectrum S) (hx : IsClosed ({x} : Set (PrimeSpectrum S))) : + IsClosed ({comap f x} : Set (PrimeSpectrum R)) := by + simpa using isClosedMap_comap_of_isIntegral f hf _ hx + +lemma closure_image_comap_zeroLocus (I : Ideal S) : + closure (comap f '' zeroLocus I) = zeroLocus (I.comap f) := by + apply subset_antisymm + · rw [(isClosed_zeroLocus _).closure_subset_iff, Set.image_subset_iff, preimage_comap_zeroLocus] + exact zeroLocus_anti_mono (Set.image_preimage_subset _ _) + · rintro x (hx : I.comap f ≤ x.asIdeal) + obtain ⟨q, hq₁, hq₂⟩ := Ideal.exists_minimalPrimes_le hx + obtain ⟨p', hp', hp'', rfl⟩ := Ideal.exists_comap_eq_of_mem_minimalPrimes f _ hq₁ + let p'' : PrimeSpectrum S := ⟨p', hp'⟩ + apply isClosed_closure.stableUnderSpecialization ((le_iff_specializes + (comap f ⟨p', hp'⟩) x).mp hq₂) (subset_closure (by exact ⟨_, hp'', rfl⟩)) + +lemma isIntegral_of_isClosedMap_comap_mapRingHom (h : IsClosedMap (comap (mapRingHom f))) : + f.IsIntegral := by + algebraize [f] + suffices Algebra.IsIntegral R S by rwa [Algebra.isIntegral_def] at this + nontriviality R + nontriviality S + constructor + intro r + let p : S[X] := C r * X - 1 + have : (1 : R[X]) ∈ Ideal.span {X} ⊔ (Ideal.span {p}).comap (mapRingHom f) := by + have H := h _ (isClosed_zeroLocus {p}) + rw [← zeroLocus_span, ← closure_eq_iff_isClosed, closure_image_comap_zeroLocus] at H + rw [← Ideal.eq_top_iff_one, sup_comm, ← zeroLocus_empty_iff_eq_top, zeroLocus_sup, H] + suffices ∀ (a : PrimeSpectrum S[X]), p ∈ a.asIdeal → X ∉ a.asIdeal by + simpa [Set.eq_empty_iff_forall_not_mem] + intro q hpq hXq + have : 1 ∈ q.asIdeal := by simpa [p] using (sub_mem (q.asIdeal.mul_mem_left (C r) hXq) hpq) + exact q.2.ne_top (q.asIdeal.eq_top_iff_one.mpr this) + obtain ⟨a, b, hb, e⟩ := Ideal.mem_span_singleton_sup.mp this + obtain ⟨c, hc : b.map (algebraMap R S) = _⟩ := Ideal.mem_span_singleton.mp hb + refine ⟨b.reverse * X ^ (1 + c.natDegree), ?_, ?_⟩ + · refine Monic.mul ?_ (by simp) + have h : b.coeff 0 = 1 := by simpa using congr(($e).coeff 0) + have : b.natTrailingDegree = 0 := by simp [h] + rw [Monic.def, reverse_leadingCoeff, trailingCoeff, this, h] + · have : p.natDegree ≤ 1 := by simpa using natDegree_linear_le (a := r) (b := -1) + rw [eval₂_eq_eval_map, reverse, Polynomial.map_mul, ← reflect_map, Polynomial.map_pow, + map_X, ← revAt_zero (1 + _), ← reflect_monomial, + ← reflect_mul _ _ (natDegree_map_le _ _) (by simp), pow_zero, mul_one, hc, + ← add_assoc, reflect_mul _ _ (this.trans (by simp)) le_rfl, + eval_mul, reflect_sub, reflect_mul _ _ (by simp) (by simp)] + simp [← pow_succ'] + +end IsIntegral + section LocalizationAtMinimal variable {I : Ideal R} [hI : I.IsPrime] From d6d71136020054d771903093d0600839425127b8 Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 19 Nov 2024 13:00:52 +0000 Subject: [PATCH 111/829] fix: replace backticks with single quotes in maintainer merge action (#19228) This should prevent mangling the `spoiler` message on Zulip. --- scripts/maintainer_merge_message.sh | 7 ++++++- 1 file changed, 6 insertions(+), 1 deletion(-) diff --git a/scripts/maintainer_merge_message.sh b/scripts/maintainer_merge_message.sh index db5a50caa1676..ec4e14808bf86 100755 --- a/scripts/maintainer_merge_message.sh +++ b/scripts/maintainer_merge_message.sh @@ -36,5 +36,10 @@ esac >&2 echo "EVENT_NAME: '${EVENT_NAME}'" >&2 printf 'COMMENT\n%s\nEND COMMENT\n' "${PR_COMMENT}" +# replace backticks in the title with single quotes +unbacktickedTitle="${PR_TITLE//\`/\'}" + +>&2 echo "neat title: '${unbacktickedTitle}'" + printf '%s requested a maintainer **%s** from %s on PR [#%s](%s):\n' "${AUTHOR}" "${M_or_D}" "${SOURCE}" "${PR}" "${URL}" -printf '```spoiler %s\n%s\n```\n' "${PR_TITLE}" "${PR_COMMENT}" +printf '```spoiler %s\n%s\n```\n' "${unbacktickedTitle}" "${PR_COMMENT}" From 62b241e911f53062101c73e0d8ec9a31076c0916 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Tue, 19 Nov 2024 15:12:10 +0000 Subject: [PATCH 112/829] feat(AlgebraicGeometry): define integral morphisms (#19121) also shows that finite = integral + locally of finite type and closed immersion = finte + mono Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib.lean | 1 + .../AlgebraicGeometry/Morphisms/Finite.lean | 60 +++++++++++++---- .../AlgebraicGeometry/Morphisms/Integral.lean | 64 +++++++++++++++++++ Mathlib/RingTheory/RingHom/Integral.lean | 32 +++++++++- 4 files changed, 143 insertions(+), 14 deletions(-) create mode 100644 Mathlib/AlgebraicGeometry/Morphisms/Integral.lean diff --git a/Mathlib.lean b/Mathlib.lean index 7566572fa39e6..49dd52b773d65 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -938,6 +938,7 @@ import Mathlib.AlgebraicGeometry.Morphisms.Finite import Mathlib.AlgebraicGeometry.Morphisms.FinitePresentation import Mathlib.AlgebraicGeometry.Morphisms.FiniteType import Mathlib.AlgebraicGeometry.Morphisms.Immersion +import Mathlib.AlgebraicGeometry.Morphisms.Integral import Mathlib.AlgebraicGeometry.Morphisms.IsIso import Mathlib.AlgebraicGeometry.Morphisms.OpenImmersion import Mathlib.AlgebraicGeometry.Morphisms.Preimmersion diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Finite.lean b/Mathlib/AlgebraicGeometry/Morphisms/Finite.lean index 3a1aca10b2cbd..0f054e90a059e 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Finite.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Finite.lean @@ -3,9 +3,8 @@ Copyright (c) 2024 Christian Merten. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Christian Merten -/ -import Mathlib.AlgebraicGeometry.Morphisms.AffineAnd -import Mathlib.AlgebraicGeometry.Morphisms.FiniteType -import Mathlib.RingTheory.RingHom.Finite +import Mathlib.AlgebraicGeometry.Morphisms.Integral +import Mathlib.Algebra.Category.Ring.Epi /-! @@ -65,15 +64,52 @@ instance (priority := 900) [IsIso f] : IsFinite f := of_isIso @IsFinite f instance {Z : Scheme.{u}} (g : Y ⟶ Z) [IsFinite f] [IsFinite g] : IsFinite (f ≫ g) := IsStableUnderComposition.comp_mem f g ‹IsFinite f› ‹IsFinite g› -instance (priority := 900) [hf : IsFinite f] : LocallyOfFiniteType f := by - have : targetAffineLocally (affineAnd @RingHom.FiniteType) f := by - rw [HasAffineProperty.eq_targetAffineLocally (P := @IsFinite)] at hf - apply targetAffineLocally_affineAnd_le RingHom.Finite.finiteType - exact hf - rw [targetAffineLocally_affineAnd_eq_affineLocally - (HasRingHomProperty.isLocal_ringHomProperty @LocallyOfFiniteType)] at this - rw [HasRingHomProperty.eq_affineLocally (P := @LocallyOfFiniteType)] - exact this.2 +lemma iff_isIntegralHom_and_locallyOfFiniteType : + IsFinite f ↔ IsIntegralHom f ∧ LocallyOfFiniteType f := by + wlog hY : IsAffine Y + · rw [IsLocalAtTarget.iff_of_openCover (P := @IsFinite) Y.affineCover, + IsLocalAtTarget.iff_of_openCover (P := @IsIntegralHom) Y.affineCover, + IsLocalAtTarget.iff_of_openCover (P := @LocallyOfFiniteType) Y.affineCover] + simp_rw [this, forall_and] + rw [HasAffineProperty.iff_of_isAffine (P := @IsFinite), + HasAffineProperty.iff_of_isAffine (P := @IsIntegralHom), + RingHom.finite_iff_isIntegral_and_finiteType, ← and_assoc] + refine and_congr_right fun ⟨_, _⟩ ↦ + (HasRingHomProperty.iff_of_isAffine (P := @LocallyOfFiniteType)).symm + +lemma eq_inf : + @IsFinite = (@IsIntegralHom ⊓ @LocallyOfFiniteType : MorphismProperty Scheme) := by + ext; exact IsFinite.iff_isIntegralHom_and_locallyOfFiniteType _ + +instance (priority := 900) [IsFinite f] : IsIntegralHom f := + ((IsFinite.iff_isIntegralHom_and_locallyOfFiniteType f).mp ‹_›).1 + +instance (priority := 900) [hf : IsFinite f] : LocallyOfFiniteType f := + ((IsFinite.iff_isIntegralHom_and_locallyOfFiniteType f).mp ‹_›).2 + +lemma _root_.AlgebraicGeometry.IsClosedImmersion.iff_isFinite_and_mono : + IsClosedImmersion f ↔ IsFinite f ∧ Mono f := by + wlog hY : IsAffine Y + · show _ ↔ _ ∧ monomorphisms _ f + rw [IsLocalAtTarget.iff_of_openCover (P := @IsFinite) Y.affineCover, + IsLocalAtTarget.iff_of_openCover (P := @IsClosedImmersion) Y.affineCover, + IsLocalAtTarget.iff_of_openCover (P := monomorphisms _) Y.affineCover] + simp_rw [this, forall_and, monomorphisms] + rw [HasAffineProperty.iff_of_isAffine (P := @IsClosedImmersion), + HasAffineProperty.iff_of_isAffine (P := @IsFinite), + RingHom.surjective_iff_epi_and_finite, @and_comm (Epi _), ← and_assoc] + refine and_congr_right fun ⟨_, _⟩ ↦ + Iff.trans ?_ (arrow_mk_iso_iff (monomorphisms _) (arrowIsoSpecΓOfIsAffine f).symm) + trans Mono (f.app ⊤).op + · exact ⟨fun h ↦ inferInstance, fun h ↦ show Epi (f.app ⊤).op.unop by infer_instance⟩ + exact (Functor.mono_map_iff_mono Scheme.Spec _).symm + +lemma _root_.AlgebraicGeometry.IsClosedImmersion.eq_isFinite_inf_mono : + @IsClosedImmersion = (@IsFinite ⊓ monomorphisms Scheme : MorphismProperty _) := by + ext; exact IsClosedImmersion.iff_isFinite_and_mono _ + +instance (priority := 900) {X Y : Scheme} (f : X ⟶ Y) [IsClosedImmersion f] : IsFinite f := + ((IsClosedImmersion.iff_isFinite_and_mono f).mp ‹_›).1 end IsFinite diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Integral.lean b/Mathlib/AlgebraicGeometry/Morphisms/Integral.lean new file mode 100644 index 0000000000000..09ee678d31470 --- /dev/null +++ b/Mathlib/AlgebraicGeometry/Morphisms/Integral.lean @@ -0,0 +1,64 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.AlgebraicGeometry.Morphisms.AffineAnd +import Mathlib.AlgebraicGeometry.Morphisms.ClosedImmersion +import Mathlib.RingTheory.RingHom.Integral + +/-! + +# Integral morphisms of schemes + +A morphism of schemes `f : X ⟶ Y` is integral if the preimage +of an arbitrary affine open subset of `Y` is affine and the induced ring map is integral. + +It is equivalent to ask only that `Y` is covered by affine opens whose preimage is affine +and the induced ring map is integral. + +## TODO + +- Show integral = universally closed + affine + +-/ + +universe v u + +open CategoryTheory TopologicalSpace Opposite MorphismProperty + +namespace AlgebraicGeometry + +/-- A morphism of schemes `X ⟶ Y` is finite if +the preimage of any affine open subset of `Y` is affine and the induced ring +hom is finite. -/ +@[mk_iff] +class IsIntegralHom {X Y : Scheme} (f : X ⟶ Y) extends IsAffineHom f : Prop where + integral_app (U : Y.Opens) (hU : IsAffineOpen U) : (f.app U).IsIntegral + +namespace IsIntegralHom + +instance hasAffineProperty : HasAffineProperty @IsIntegralHom + fun X _ f _ ↦ IsAffine X ∧ RingHom.IsIntegral (f.app ⊤) := by + show HasAffineProperty @IsIntegralHom (affineAnd RingHom.IsIntegral) + rw [HasAffineProperty.affineAnd_iff _ RingHom.isIntegral_respectsIso + RingHom.isIntegral_isStableUnderBaseChange.localizationPreserves + RingHom.isIntegral_ofLocalizationSpan] + simp [isIntegralHom_iff] + +instance : IsStableUnderComposition @IsIntegralHom := + HasAffineProperty.affineAnd_isStableUnderComposition (Q := RingHom.IsIntegral) hasAffineProperty + RingHom.isIntegral_stableUnderComposition + +instance : IsStableUnderBaseChange @IsIntegralHom := + HasAffineProperty.affineAnd_isStableUnderBaseChange (Q := RingHom.IsIntegral) hasAffineProperty + RingHom.isIntegral_respectsIso RingHom.isIntegral_isStableUnderBaseChange + +instance : ContainsIdentities @IsIntegralHom := + ⟨fun X ↦ ⟨fun _ _ ↦ by simpa using RingHom.isIntegral_of_surjective _ (Equiv.refl _).surjective⟩⟩ + +instance : IsMultiplicative @IsIntegralHom where + +end IsIntegralHom + +end AlgebraicGeometry diff --git a/Mathlib/RingTheory/RingHom/Integral.lean b/Mathlib/RingTheory/RingHom/Integral.lean index 62662762e06d7..2662ea9b1ccba 100644 --- a/Mathlib/RingTheory/RingHom/Integral.lean +++ b/Mathlib/RingTheory/RingHom/Integral.lean @@ -3,8 +3,8 @@ Copyright (c) 2021 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ -import Mathlib.RingTheory.RingHomProperties -import Mathlib.RingTheory.IntegralClosure.IsIntegralClosure.Basic +import Mathlib.RingTheory.LocalProperties.Basic +import Mathlib.RingTheory.Localization.Integral /-! @@ -36,4 +36,32 @@ theorem isIntegral_isStableUnderBaseChange : IsStableUnderBaseChange fun f => f. · intro x y; exact IsIntegral.tmul x (h y) · intro x y hx hy; exact IsIntegral.add hx hy +open Polynomial in +/-- `S` is an integral `R`-algebra if there exists a set `{ r }` that + spans `R` such that each `Sᵣ` is an integral `Rᵣ`-algebra. -/ +theorem isIntegral_ofLocalizationSpan : + OfLocalizationSpan (RingHom.IsIntegral ·) := by + introv R hs H r + letI := f.toAlgebra + show r ∈ (integralClosure R S).toSubmodule + apply Submodule.mem_of_span_eq_top_of_smul_pow_mem _ s hs + rintro ⟨t, ht⟩ + letI := (Localization.awayMap f t).toAlgebra + haveI : IsScalarTower R (Localization.Away t) (Localization.Away (f t)) := .of_algebraMap_eq' + (IsLocalization.lift_comp _).symm + have : _root_.IsIntegral (Localization.Away t) (algebraMap S (Localization.Away (f t)) r) := + H ⟨t, ht⟩ (algebraMap _ _ r) + obtain ⟨⟨_, n, rfl⟩, p, hp, hp'⟩ := this.exists_multiple_integral_of_isLocalization (.powers t) + rw [IsScalarTower.algebraMap_eq R S, Submonoid.smul_def, Algebra.smul_def, + IsScalarTower.algebraMap_apply R S, ← map_mul, ← hom_eval₂, + IsLocalization.map_eq_zero_iff (.powers (f t))] at hp' + obtain ⟨⟨x, m, (rfl : algebraMap R S t ^ m = x)⟩, e⟩ := hp' + by_cases hp' : 1 ≤ p.natDegree; swap + · obtain rfl : p = 1 := eq_one_of_monic_natDegree_zero hp (by nlinarith) + exact ⟨m, by simp [Algebra.smul_def, show algebraMap R S t ^ m = 0 by simpa using e]⟩ + refine ⟨m + n, p.scaleRoots (t ^ m), (monic_scaleRoots_iff _).mpr hp, ?_⟩ + have := p.scaleRoots_eval₂_mul (algebraMap R S) (t ^ n • r) (t ^ m) + simp only [pow_add, ← Algebra.smul_def, mul_smul, ← map_pow] at e this ⊢ + rw [this, ← tsub_add_cancel_of_le hp', pow_succ, mul_smul, e, smul_zero] + end RingHom From e948d859635b67a2d40b53d10bcf906226e2a460 Mon Sep 17 00:00:00 2001 From: grunweg Date: Tue, 19 Nov 2024 15:27:44 +0000 Subject: [PATCH 113/829] chore: further renames (#19243) Similar to #19209: bring more lemma names in line with the naming convention. --- .../PrimeSpectrum/Basic.lean | 2 +- Mathlib/AlgebraicGeometry/Properties.lean | 2 +- Mathlib/Dynamics/Minimal.lean | 4 +++- Mathlib/MeasureTheory/Group/Action.lean | 2 +- Mathlib/MeasureTheory/Group/Measure.lean | 4 ++-- Mathlib/MeasureTheory/Measure/Regular.lean | 6 ++++-- Mathlib/Topology/ContinuousMap/Algebra.lean | 10 +++++----- Mathlib/Topology/Irreducible.lean | 9 +++++++-- Mathlib/Topology/Maps/Basic.lean | 10 ++++++---- Mathlib/Topology/NoetherianSpace.lean | 6 +++--- .../UniformSpace/CompactConvergence.lean | 19 +++++++++++-------- 11 files changed, 44 insertions(+), 30 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index b1fe379b4c1cc..8121592d7b482 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -156,7 +156,7 @@ theorem isIrreducible_zeroLocus_iff_of_radical (I : Ideal R) (hI : I.IsRadical) apply and_congr · rw [Set.nonempty_iff_ne_empty, Ne, zeroLocus_empty_iff_eq_top] · trans ∀ x y : Ideal R, Z(I) ⊆ Z(x) ∪ Z(y) → Z(I) ⊆ Z(x) ∨ Z(I) ⊆ Z(y) - · simp_rw [isPreirreducible_iff_closed_union_closed, isClosed_iff_zeroLocus_ideal] + · simp_rw [isPreirreducible_iff_isClosed_union_isClosed, isClosed_iff_zeroLocus_ideal] constructor · rintro h x y exact h _ _ ⟨x, rfl⟩ ⟨y, rfl⟩ diff --git a/Mathlib/AlgebraicGeometry/Properties.lean b/Mathlib/AlgebraicGeometry/Properties.lean index 800d80f953765..79091bf28dd7a 100644 --- a/Mathlib/AlgebraicGeometry/Properties.lean +++ b/Mathlib/AlgebraicGeometry/Properties.lean @@ -203,7 +203,7 @@ instance irreducibleSpace_of_isIntegral [IsIntegral X] : IrreducibleSpace X := b replace H : ¬IsPreirreducible (⊤ : Set X) := fun h => H { toPreirreducibleSpace := ⟨h⟩ toNonempty := inferInstance } - simp_rw [isPreirreducible_iff_closed_union_closed, not_forall, not_or] at H + simp_rw [isPreirreducible_iff_isClosed_union_isClosed, not_forall, not_or] at H rcases H with ⟨S, T, hS, hT, h₁, h₂, h₃⟩ erw [not_forall] at h₂ h₃ simp_rw [not_forall] at h₂ h₃ diff --git a/Mathlib/Dynamics/Minimal.lean b/Mathlib/Dynamics/Minimal.lean index ce019b475eaf4..c29a7f2fce92e 100644 --- a/Mathlib/Dynamics/Minimal.lean +++ b/Mathlib/Dynamics/Minimal.lean @@ -93,7 +93,7 @@ theorem eq_empty_or_univ_of_smul_invariant_closed [IsMinimal M α] {s : Set α} hs.closure_eq ▸ (dense_of_nonempty_smul_invariant M hne hsmul).closure_eq @[to_additive] -theorem isMinimal_iff_closed_smul_invariant [ContinuousConstSMul M α] : +theorem isMinimal_iff_isClosed_smul_invariant [ContinuousConstSMul M α] : IsMinimal M α ↔ ∀ s : Set α, IsClosed s → (∀ c : M, c • s ⊆ s) → s = ∅ ∨ s = univ := by constructor · intro _ _ @@ -101,3 +101,5 @@ theorem isMinimal_iff_closed_smul_invariant [ContinuousConstSMul M α] : refine fun H ↦ ⟨fun _ ↦ dense_iff_closure_eq.2 <| (H _ ?_ ?_).resolve_left ?_⟩ exacts [isClosed_closure, fun _ ↦ smul_closure_orbit_subset _ _, (orbit_nonempty _).closure.ne_empty] +@[deprecated (since := "2024-11-19")] alias +isMinimal_iff_closed_smul_invariant := isMinimal_iff_isClosed_smul_invariant diff --git a/Mathlib/MeasureTheory/Group/Action.lean b/Mathlib/MeasureTheory/Group/Action.lean index 04f432bce67b9..0cb212b5273b1 100644 --- a/Mathlib/MeasureTheory/Group/Action.lean +++ b/Mathlib/MeasureTheory/Group/Action.lean @@ -307,7 +307,7 @@ variable [Measure.Regular μ] @[to_additive] theorem measure_isOpen_pos_of_smulInvariant_of_ne_zero (hμ : μ ≠ 0) (hU : IsOpen U) (hne : U.Nonempty) : 0 < μ U := - let ⟨_K, hK, hμK⟩ := Regular.exists_compact_not_null.mpr hμ + let ⟨_K, hK, hμK⟩ := Regular.exists_isCompact_not_null.mpr hμ measure_isOpen_pos_of_smulInvariant_of_compact_ne_zero G hK hμK hU hne @[to_additive] diff --git a/Mathlib/MeasureTheory/Group/Measure.lean b/Mathlib/MeasureTheory/Group/Measure.lean index 05604e5fa4945..8662ff3b7f53c 100644 --- a/Mathlib/MeasureTheory/Group/Measure.lean +++ b/Mathlib/MeasureTheory/Group/Measure.lean @@ -502,7 +502,7 @@ theorem isOpenPosMeasure_of_mulLeftInvariant_of_compact (K : Set G) (hK : IsComp @[to_additive "A nonzero left-invariant regular measure gives positive mass to any open set."] instance (priority := 80) isOpenPosMeasure_of_mulLeftInvariant_of_regular [Regular μ] [NeZero μ] : IsOpenPosMeasure μ := - let ⟨K, hK, h2K⟩ := Regular.exists_compact_not_null.mpr (NeZero.ne μ) + let ⟨K, hK, h2K⟩ := Regular.exists_isCompact_not_null.mpr (NeZero.ne μ) isOpenPosMeasure_of_mulLeftInvariant_of_compact K hK h2K /-- A nonzero left-invariant inner regular measure gives positive mass to any open set. -/ @@ -510,7 +510,7 @@ instance (priority := 80) isOpenPosMeasure_of_mulLeftInvariant_of_regular [Regul instance (priority := 80) isOpenPosMeasure_of_mulLeftInvariant_of_innerRegular [InnerRegular μ] [NeZero μ] : IsOpenPosMeasure μ := - let ⟨K, hK, h2K⟩ := InnerRegular.exists_compact_not_null.mpr (NeZero.ne μ) + let ⟨K, hK, h2K⟩ := InnerRegular.exists_isCompact_not_null.mpr (NeZero.ne μ) isOpenPosMeasure_of_mulLeftInvariant_of_compact K hK h2K @[to_additive] diff --git a/Mathlib/MeasureTheory/Measure/Regular.lean b/Mathlib/MeasureTheory/Measure/Regular.lean index 00bbeafec255e..ed27aa6936522 100644 --- a/Mathlib/MeasureTheory/Measure/Regular.lean +++ b/Mathlib/MeasureTheory/Measure/Regular.lean @@ -663,9 +663,10 @@ lemma innerRegularWRT_isClosed_isOpen [R1Space α] [OpensMeasurableSpace α] [h exact ⟨closure K, K_comp.closure_subset_of_isOpen hU KU, isClosed_closure, hK.trans_le (measure_mono subset_closure)⟩ -theorem exists_compact_not_null [InnerRegular μ] : (∃ K, IsCompact K ∧ μ K ≠ 0) ↔ μ ≠ 0 := by +theorem exists_isCompact_not_null [InnerRegular μ] : (∃ K, IsCompact K ∧ μ K ≠ 0) ↔ μ ≠ 0 := by simp_rw [Ne, ← measure_univ_eq_zero, MeasurableSet.univ.measure_eq_iSup_isCompact, ENNReal.iSup_eq_zero, not_forall, exists_prop, subset_univ, true_and] +@[deprecated (since := "2024-11-19")] alias exists_compact_not_null := exists_isCompact_not_null /-- If `μ` is inner regular, then any measurable set can be approximated by a compact subset. See also `MeasurableSet.exists_isCompact_lt_add_of_ne_top`. -/ @@ -986,9 +987,10 @@ theorem _root_.IsOpen.measure_eq_iSup_isCompact ⦃U : Set α⦄ (hU : IsOpen U) [Regular μ] : μ U = ⨆ (K : Set α) (_ : K ⊆ U) (_ : IsCompact K), μ K := Regular.innerRegular.measure_eq_iSup hU -theorem exists_compact_not_null [Regular μ] : (∃ K, IsCompact K ∧ μ K ≠ 0) ↔ μ ≠ 0 := by +theorem exists_isCompact_not_null [Regular μ] : (∃ K, IsCompact K ∧ μ K ≠ 0) ↔ μ ≠ 0 := by simp_rw [Ne, ← measure_univ_eq_zero, isOpen_univ.measure_eq_iSup_isCompact, ENNReal.iSup_eq_zero, not_forall, exists_prop, subset_univ, true_and] +@[deprecated (since := "2024-11-19")] alias exists_compact_not_null := exists_isCompact_not_null /-- If `μ` is a regular measure, then any measurable set of finite measure can be approximated by a compact subset. See also `MeasurableSet.exists_isCompact_lt_add` and diff --git a/Mathlib/Topology/ContinuousMap/Algebra.lean b/Mathlib/Topology/ContinuousMap/Algebra.lean index 2cefc72b970be..42a37d593e9a0 100644 --- a/Mathlib/Topology/ContinuousMap/Algebra.lean +++ b/Mathlib/Topology/ContinuousMap/Algebra.lean @@ -351,20 +351,20 @@ instance [CommGroup β] [TopologicalGroup β] : TopologicalGroup C(α, β) where have : UniformGroup β := comm_topologicalGroup_is_uniform rw [continuous_iff_continuousAt] rintro ⟨f, g⟩ - rw [ContinuousAt, tendsto_iff_forall_compact_tendstoUniformlyOn, nhds_prod_eq] + rw [ContinuousAt, tendsto_iff_forall_isCompact_tendstoUniformlyOn, nhds_prod_eq] exact fun K hK => uniformContinuous_mul.comp_tendstoUniformlyOn - ((tendsto_iff_forall_compact_tendstoUniformlyOn.mp Filter.tendsto_id K hK).prod - (tendsto_iff_forall_compact_tendstoUniformlyOn.mp Filter.tendsto_id K hK)) + ((tendsto_iff_forall_isCompact_tendstoUniformlyOn.mp Filter.tendsto_id K hK).prod + (tendsto_iff_forall_isCompact_tendstoUniformlyOn.mp Filter.tendsto_id K hK)) continuous_inv := by letI : UniformSpace β := TopologicalGroup.toUniformSpace β have : UniformGroup β := comm_topologicalGroup_is_uniform rw [continuous_iff_continuousAt] intro f - rw [ContinuousAt, tendsto_iff_forall_compact_tendstoUniformlyOn] + rw [ContinuousAt, tendsto_iff_forall_isCompact_tendstoUniformlyOn] exact fun K hK => uniformContinuous_inv.comp_tendstoUniformlyOn - (tendsto_iff_forall_compact_tendstoUniformlyOn.mp Filter.tendsto_id K hK) + (tendsto_iff_forall_isCompact_tendstoUniformlyOn.mp Filter.tendsto_id K hK) /-- If an infinite product of functions in `C(α, β)` converges to `g` (for the compact-open topology), then the pointwise product converges to `g x` for all `x ∈ α`. -/ diff --git a/Mathlib/Topology/Irreducible.lean b/Mathlib/Topology/Irreducible.lean index c5c1339a53ef9..5d383423319f0 100644 --- a/Mathlib/Topology/Irreducible.lean +++ b/Mathlib/Topology/Irreducible.lean @@ -240,7 +240,7 @@ theorem isIrreducible_iff_sInter : /-- A set is preirreducible if and only if for every cover by two closed sets, it is contained in one of the two covering sets. -/ -theorem isPreirreducible_iff_closed_union_closed : +theorem isPreirreducible_iff_isClosed_union_isClosed : IsPreirreducible s ↔ ∀ z₁ z₂ : Set X, IsClosed z₁ → IsClosed z₂ → s ⊆ z₁ ∪ z₂ → s ⊆ z₁ ∨ s ⊆ z₂ := by refine compl_surjective.forall.trans <| forall_congr' fun z₁ => compl_surjective.forall.trans <| @@ -248,10 +248,12 @@ theorem isPreirreducible_iff_closed_union_closed : simp only [isOpen_compl_iff, ← compl_union, inter_compl_nonempty_iff] refine forall₂_congr fun _ _ => ?_ rw [← and_imp, ← not_or, not_imp_not] +@[deprecated (since := "2024-11-19")] alias +isPreirreducible_iff_closed_union_closed := isPreirreducible_iff_isClosed_union_isClosed /-- A set is irreducible if and only if for every cover by a finite collection of closed sets, it is contained in one of the members of the collection. -/ -theorem isIrreducible_iff_sUnion_closed : +theorem isIrreducible_iff_sUnion_isClosed : IsIrreducible s ↔ ∀ t : Finset (Set X), (∀ z ∈ t, IsClosed z) → (s ⊆ ⋃₀ ↑t) → ∃ z ∈ t, s ⊆ z := by simp only [isIrreducible_iff_sInter] @@ -263,6 +265,9 @@ theorem isIrreducible_iff_sUnion_closed : simp only [not_exists, not_and, ← compl_iInter₂, ← sInter_eq_biInter, subset_compl_iff_disjoint_right, not_disjoint_iff_nonempty_inter] +@[deprecated (since := "2024-11-19")] alias +isIrreducible_iff_sUnion_closed := isIrreducible_iff_sUnion_isClosed + /-- A nonempty open subset of a preirreducible subspace is dense in the subspace. -/ theorem subset_closure_inter_of_isPreirreducible_of_isOpen {S U : Set X} (hS : IsPreirreducible S) (hU : IsOpen U) (h : (S ∩ U).Nonempty) : S ⊆ closure (S ∩ U) := by diff --git a/Mathlib/Topology/Maps/Basic.lean b/Mathlib/Topology/Maps/Basic.lean index b71fbe5d98fd9..077466f86f877 100644 --- a/Mathlib/Topology/Maps/Basic.lean +++ b/Mathlib/Topology/Maps/Basic.lean @@ -286,13 +286,15 @@ lemma isQuotientMap_iff : IsQuotientMap f ↔ Surjective f ∧ ∀ s, IsOpen s @[deprecated (since := "2024-10-22")] alias quotientMap_iff := isQuotientMap_iff -theorem isQuotientMap_iff_closed : +theorem isQuotientMap_iff_isClosed : IsQuotientMap f ↔ Surjective f ∧ ∀ s : Set Y, IsClosed s ↔ IsClosed (f ⁻¹' s) := isQuotientMap_iff.trans <| Iff.rfl.and <| compl_surjective.forall.trans <| by simp only [isOpen_compl_iff, preimage_compl] @[deprecated (since := "2024-10-22")] -alias quotientMap_iff_closed := isQuotientMap_iff_closed +alias quotientMap_iff_closed := isQuotientMap_iff_isClosed +@[deprecated (since := "2024-11-19")] +alias isQuotientMap_iff_closed := isQuotientMap_iff_isClosed namespace IsQuotientMap @@ -326,7 +328,7 @@ protected lemma isOpen_preimage (hf : IsQuotientMap f) {s : Set Y} : IsOpen (f protected theorem isClosed_preimage (hf : IsQuotientMap f) {s : Set Y} : IsClosed (f ⁻¹' s) ↔ IsClosed s := - ((isQuotientMap_iff_closed.1 hf).2 s).symm + ((isQuotientMap_iff_isClosed.1 hf).2 s).symm end IsQuotientMap @@ -489,7 +491,7 @@ theorem isClosed_range (hf : IsClosedMap f) : IsClosed (range f) := theorem isQuotientMap (hcl : IsClosedMap f) (hcont : Continuous f) (hsurj : Surjective f) : IsQuotientMap f := - isQuotientMap_iff_closed.2 ⟨hsurj, fun s => + isQuotientMap_iff_isClosed.2 ⟨hsurj, fun s => ⟨fun hs => hs.preimage hcont, fun hs => hsurj.image_preimage s ▸ hcl _ hs⟩⟩ @[deprecated (since := "2024-10-22")] diff --git a/Mathlib/Topology/NoetherianSpace.lean b/Mathlib/Topology/NoetherianSpace.lean index 4008eacfe3e3f..64689fd04f3f0 100644 --- a/Mathlib/Topology/NoetherianSpace.lean +++ b/Mathlib/Topology/NoetherianSpace.lean @@ -165,7 +165,7 @@ theorem NoetherianSpace.exists_finite_set_closeds_irreducible [NoetherianSpace · by_cases h₁ : IsPreirreducible (s : Set α) · replace h₁ : IsIrreducible (s : Set α) := ⟨Closeds.coe_nonempty.2 h₀, h₁⟩ use {s}; simp [h₁] - · simp only [isPreirreducible_iff_closed_union_closed, not_forall, not_or] at h₁ + · simp only [isPreirreducible_iff_isClosed_union_isClosed, not_forall, not_or] at h₁ obtain ⟨z₁, z₂, hz₁, hz₂, h, hz₁', hz₂'⟩ := h₁ lift z₁ to Closeds α using hz₁ lift z₂ to Closeds α using hz₂ @@ -198,7 +198,7 @@ theorem NoetherianSpace.finite_irreducibleComponents [NoetherianSpace α] : NoetherianSpace.exists_finite_set_isClosed_irreducible isClosed_univ (α := α) refine hSf.subset fun s hs => ?_ lift S to Finset (Set α) using hSf - rcases isIrreducible_iff_sUnion_closed.1 hs.1 S hSc (hSU ▸ Set.subset_univ _) with ⟨t, htS, ht⟩ + rcases isIrreducible_iff_sUnion_isClosed.1 hs.1 S hSc (hSU ▸ Set.subset_univ _) with ⟨t, htS, ht⟩ rwa [ht.antisymm (hs.2 (hSi _ htS) ht)] /-- [Stacks: Lemma 0052 (3)](https://stacks.math.columbia.edu/tag/0052) -/ @@ -213,7 +213,7 @@ theorem NoetherianSpace.exists_open_ne_empty_le_irreducibleComponent [Noetherian let U := Z \ ⋃ (x : ι), x have hU0 : U ≠ ∅ := fun r ↦ by - obtain ⟨Z', hZ'⟩ := isIrreducible_iff_sUnion_closed.mp H.1 hι.toFinset + obtain ⟨Z', hZ'⟩ := isIrreducible_iff_sUnion_isClosed.mp H.1 hι.toFinset (fun z hz ↦ by simp only [Set.Finite.mem_toFinset, Set.mem_diff, Set.mem_singleton_iff] at hz exact isClosed_of_mem_irreducibleComponents _ hz.1) diff --git a/Mathlib/Topology/UniformSpace/CompactConvergence.lean b/Mathlib/Topology/UniformSpace/CompactConvergence.lean index 0a9a0c7520df8..e39e65d2fbda4 100644 --- a/Mathlib/Topology/UniformSpace/CompactConvergence.lean +++ b/Mathlib/Topology/UniformSpace/CompactConvergence.lean @@ -48,12 +48,12 @@ and also prove its basic properties. a similar statement that uses a basis of entourages of `β` instead of all entourages. It is useful, e.g., if `β` is a metric space. -* `ContinuousMap.tendsto_iff_forall_compact_tendstoUniformlyOn`: +* `ContinuousMap.tendsto_iff_forall_isCompact_tendstoUniformlyOn`: a sequence of functions `Fₙ` in `C(α, β)` converges in the compact-open topology to some `f` iff `Fₙ` converges to `f` uniformly on each compact subset `K` of `α`. * Topology induced by the uniformity described above agrees with the compact-open topology. - This is essentially the same as `ContinuousMap.tendsto_iff_forall_compact_tendstoUniformlyOn`. + This is essentially the same as `ContinuousMap.tendsto_iff_forall_isCompact_tendstoUniformlyOn`. This fact is not available as a separate theorem. Instead, we override the projection of `ContinuousMap.compactConvergenceUniformity` @@ -95,7 +95,7 @@ namespace ContinuousMap /-- Compact-open topology on `C(α, β)` agrees with the topology of uniform convergence on compacts: a family of continuous functions `F i` tends to `f` in the compact-open topology if and only if the `F i` tends to `f` uniformly on all compact sets. -/ -theorem tendsto_iff_forall_compact_tendstoUniformlyOn +theorem tendsto_iff_forall_isCompact_tendstoUniformlyOn {ι : Type u₃} {p : Filter ι} {F : ι → C(α, β)} {f} : Tendsto F p (𝓝 f) ↔ ∀ K, IsCompact K → TendstoUniformlyOn (fun i a => F i a) f p K := by rw [tendsto_nhds_compactOpen] @@ -138,6 +138,9 @@ theorem tendsto_iff_forall_compact_tendstoUniformlyOn -- maps `K` to `U` as well filter_upwards [h K hK V hV] with g hg x hx using hVf _ (mem_image_of_mem f hx) (hg x hx) +@[deprecated (since := "2024-11-19")] alias +tendsto_iff_forall_compact_tendstoUniformlyOn := tendsto_iff_forall_isCompact_tendstoUniformlyOn + /-- Interpret a bundled continuous map as an element of `α →ᵤ[{K | IsCompact K}] β`. We use this map to induce the `UniformSpace` structure on `C(α, β)`. -/ @@ -157,14 +160,14 @@ open UniformSpace in The uniformity comes from `α →ᵤ[{K | IsCompact K}] β` (i.e., `UniformOnFun α β {K | IsCompact K}`) which defines topology of uniform convergence on compact sets. -We use `ContinuousMap.tendsto_iff_forall_compact_tendstoUniformlyOn` +We use `ContinuousMap.tendsto_iff_forall_isCompact_tendstoUniformlyOn` to show that the induced topology agrees with the compact-open topology and replace the topology with `compactOpen` to avoid non-defeq diamonds, see Note [forgetful inheritance]. -/ instance compactConvergenceUniformSpace : UniformSpace C(α, β) := .replaceTopology (.comap toUniformOnFunIsCompact inferInstance) <| by refine TopologicalSpace.ext_nhds fun f ↦ eq_of_forall_le_iff fun l ↦ ?_ - simp_rw [← tendsto_id', tendsto_iff_forall_compact_tendstoUniformlyOn, + simp_rw [← tendsto_id', tendsto_iff_forall_isCompact_tendstoUniformlyOn, nhds_induced, tendsto_comap_iff, UniformOnFun.tendsto_iff_tendstoUniformlyOn] rfl @@ -232,7 +235,7 @@ variable {ι : Type u₃} {p : Filter ι} {F : ι → C(α, β)} {f} /-- Locally uniform convergence implies convergence in the compact-open topology. -/ theorem tendsto_of_tendstoLocallyUniformly (h : TendstoLocallyUniformly (fun i a => F i a) f p) : Tendsto F p (𝓝 f) := by - rw [tendsto_iff_forall_compact_tendstoUniformlyOn] + rw [tendsto_iff_forall_isCompact_tendstoUniformlyOn] intro K hK rw [← tendstoLocallyUniformlyOn_iff_tendstoUniformlyOn_of_compact hK] exact h.tendstoLocallyUniformlyOn @@ -245,7 +248,7 @@ see `ContinuousMap.tendsto_of_tendstoLocallyUniformly`. -/ theorem tendsto_iff_tendstoLocallyUniformly [WeaklyLocallyCompactSpace α] : Tendsto F p (𝓝 f) ↔ TendstoLocallyUniformly (fun i a => F i a) f p := by refine ⟨fun h V hV x ↦ ?_, tendsto_of_tendstoLocallyUniformly⟩ - rw [tendsto_iff_forall_compact_tendstoUniformlyOn] at h + rw [tendsto_iff_forall_isCompact_tendstoUniformlyOn] at h obtain ⟨n, hn₁, hn₂⟩ := exists_compact_mem_nhds x exact ⟨n, hn₂, h n hn₁ V hV⟩ @@ -313,7 +316,7 @@ theorem hasBasis_compactConvergenceUniformity_of_compact : continuous functions on a compact space. -/ theorem tendsto_iff_tendstoUniformly : Tendsto F p (𝓝 f) ↔ TendstoUniformly (fun i a => F i a) f p := by - rw [tendsto_iff_forall_compact_tendstoUniformlyOn, ← tendstoUniformlyOn_univ] + rw [tendsto_iff_forall_isCompact_tendstoUniformlyOn, ← tendstoUniformlyOn_univ] exact ⟨fun h => h univ isCompact_univ, fun h K _hK => h.mono (subset_univ K)⟩ end CompactDomain From fa5742c52990769eba4704185c51c630902c1fa4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 19 Nov 2024 15:36:28 +0000 Subject: [PATCH 114/829] =?UTF-8?q?feat:=20`a=20^=20m=20=E2=89=A4=20b=20^?= =?UTF-8?q?=20n`=20as=20a=20`gcongr`=20lemma=20(#19220)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From GrowthInGroups --- Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean | 4 ++++ MathlibTest/GCongr/inequalities.lean | 2 +- 2 files changed, 5 insertions(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean index 274bb53619f6f..0a60e09d11d81 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean @@ -163,6 +163,10 @@ theorem Monotone.pow_const {f : β → M} (hf : Monotone f) : ∀ n : ℕ, Monot @[to_additive nsmul_right_mono] theorem pow_left_mono (n : ℕ) : Monotone fun a : M => a ^ n := monotone_id.pow_const _ +@[to_additive (attr := gcongr)] +lemma pow_le_pow {a b : M} (hab : a ≤ b) (ht : 1 ≤ b) {m n : ℕ} (hmn : m ≤ n) : a ^ m ≤ b ^ n := + (pow_le_pow_left' hab _).trans (pow_le_pow_right' ht hmn) + end CovariantLESwap end Preorder diff --git a/MathlibTest/GCongr/inequalities.lean b/MathlibTest/GCongr/inequalities.lean index 2bce82a11a237..3056ee3b5e6aa 100644 --- a/MathlibTest/GCongr/inequalities.lean +++ b/MathlibTest/GCongr/inequalities.lean @@ -100,7 +100,7 @@ example {a b : ℚ} (h₁ : a < b) (c : ℝ) : (a + c : ℝ) < b + c := by gcong example {k m n : ℤ} (H : m ^ 2 ≤ n ^ 2) : k + m ^ 2 ≤ k + n ^ 2 := by gcongr -- test of behaviour when no lemmas are applicable -example (n k : ℕ) (H : n ^ k + 1 ≤ k ^ n + 1) : n ^ k ≤ k ^ n := by +example (n k : ℕ) (H : n % k + 1 ≤ k % n + 1) : n % k ≤ k % n := by success_if_fail_with_msg "gcongr did not make progress" (gcongr) From 265b1fea02687131c71bdb29a01c8c1da4f1f6f1 Mon Sep 17 00:00:00 2001 From: peabrainiac Date: Tue, 19 Nov 2024 16:30:32 +0000 Subject: [PATCH 115/829] feat(Topology/Connected/PathConnected): some instances for locally path-connected spaces (#17064) Shows that locally path-connected spaces are locally connected, and that quotients and disjoint unions of locally path-connected spaces are locally path-connected. Co-authored-by: peabrainiac <43812953+peabrainiac@users.noreply.github.com> --- Mathlib/Topology/Connected/PathConnected.lean | 112 +++++++++++++++++- 1 file changed, 107 insertions(+), 5 deletions(-) diff --git a/Mathlib/Topology/Connected/PathConnected.lean b/Mathlib/Topology/Connected/PathConnected.lean index 07f52c0cebe24..82ff89fe8d4c4 100644 --- a/Mathlib/Topology/Connected/PathConnected.lean +++ b/Mathlib/Topology/Connected/PathConnected.lean @@ -1160,13 +1160,18 @@ alias locPathConnected_of_bases := LocPathConnectedSpace.of_bases variable [LocPathConnectedSpace X] +protected theorem IsOpen.pathComponentIn (x : X) (hF : IsOpen F) : + IsOpen (pathComponentIn x F) := by + rw [isOpen_iff_mem_nhds] + intro y hy + let ⟨s, hs⟩ := (path_connected_basis y).mem_iff.mp (hF.mem_nhds (pathComponentIn_subset hy)) + exact mem_of_superset hs.1.1 <| pathComponentIn_congr hy ▸ + hs.1.2.subset_pathComponentIn (mem_of_mem_nhds hs.1.1) hs.2 + /-- In a locally path connected space, each path component is an open set. -/ protected theorem IsOpen.pathComponent (x : X) : IsOpen (pathComponent x) := by - rw [isOpen_iff_mem_nhds] - intro y hxy - rcases (path_connected_basis y).ex_mem with ⟨V, hVy, hVc⟩ - filter_upwards [hVy] with z hz - exact hxy.out.trans (hVc.joinedIn _ (mem_of_mem_nhds hVy) _ hz).joined + rw [← pathComponentIn_univ] + exact isOpen_univ.pathComponentIn _ /-- In a locally path connected space, each path component is a closed set. -/ protected theorem IsClosed.pathComponent (x : X) : IsClosed (pathComponent x) := by @@ -1192,6 +1197,13 @@ theorem pathConnected_subset_basis {U : Set X} (h : IsOpen U) (hx : x ∈ U) : (𝓝 x).HasBasis (fun s : Set X => s ∈ 𝓝 x ∧ IsPathConnected s ∧ s ⊆ U) id := (path_connected_basis x).hasBasis_self_subset (IsOpen.mem_nhds h hx) +theorem isOpen_isPathConnected_basis (x : X) : + (𝓝 x).HasBasis (fun s : Set X ↦ IsOpen s ∧ x ∈ s ∧ IsPathConnected s) id := by + refine ⟨fun s ↦ ⟨fun hs ↦ ?_, fun ⟨u, hu⟩ ↦ mem_nhds_iff.mpr ⟨u, hu.2, hu.1.1, hu.1.2.1⟩⟩⟩ + have ⟨u, hus, hu, hxu⟩ := mem_nhds_iff.mp hs + exact ⟨pathComponentIn x u, ⟨hu.pathComponentIn _, ⟨mem_pathComponentIn_self hxu, + isPathConnected_pathComponentIn hxu⟩⟩, pathComponentIn_subset.trans hus⟩ + theorem Topology.IsOpenEmbedding.locPathConnectedSpace {e : Y → X} (he : IsOpenEmbedding e) : LocPathConnectedSpace Y := have (y : Y) : @@ -1215,4 +1227,94 @@ theorem IsOpen.isConnected_iff_isPathConnected {U : Set X} (U_op : IsOpen U) : haveI := U_op.locPathConnectedSpace exact pathConnectedSpace_iff_connectedSpace.symm +/-- Locally path-connected spaces are locally connected. -/ +instance : LocallyConnectedSpace X := by + refine ⟨forall_imp (fun x h ↦ ⟨fun s ↦ ?_⟩) isOpen_isPathConnected_basis⟩ + refine ⟨fun hs ↦ ?_, fun ⟨u, ⟨hu, hxu, _⟩, hus⟩ ↦ mem_nhds_iff.mpr ⟨u, hus, hu, hxu⟩⟩ + let ⟨u, ⟨hu, hxu, hu'⟩, hus⟩ := (h.mem_iff' s).mp hs + exact ⟨u, ⟨hu, hxu, hu'.isConnected⟩, hus⟩ + +/-- A space is locally path-connected iff all path components of open subsets are open. -/ +lemma locPathConnectedSpace_iff_isOpen_pathComponentIn {X : Type*} [TopologicalSpace X] : + LocPathConnectedSpace X ↔ ∀ (x : X) (u : Set X), IsOpen u → IsOpen (pathComponentIn x u) := + ⟨fun _ _ _ hu ↦ hu.pathComponentIn _, fun h ↦ ⟨fun x ↦ ⟨fun s ↦ by + refine ⟨fun hs ↦ ?_, fun ⟨_, ht⟩ ↦ Filter.mem_of_superset ht.1.1 ht.2⟩ + let ⟨u, hu⟩ := mem_nhds_iff.mp hs + exact ⟨pathComponentIn x u, ⟨(h x u hu.2.1).mem_nhds (mem_pathComponentIn_self hu.2.2), + isPathConnected_pathComponentIn hu.2.2⟩, pathComponentIn_subset.trans hu.1⟩⟩⟩⟩ + +/-- A space is locally path-connected iff all path components of open subsets are neighbourhoods. -/ +lemma locPathConnectedSpace_iff_pathComponentIn_mem_nhds {X : Type*} [TopologicalSpace X] : + LocPathConnectedSpace X ↔ + ∀ x : X, ∀ u : Set X, IsOpen u → x ∈ u → pathComponentIn x u ∈ nhds x := by + rw [locPathConnectedSpace_iff_isOpen_pathComponentIn] + simp_rw [forall_comm (β := Set X), ← imp_forall_iff] + refine forall_congr' fun u ↦ imp_congr_right fun _ ↦ ?_ + exact ⟨fun h x hxu ↦ (h x).mem_nhds (mem_pathComponentIn_self hxu), + fun h x ↦ isOpen_iff_mem_nhds.mpr fun y hy ↦ + pathComponentIn_congr hy ▸ h y <| pathComponentIn_subset hy⟩ + +/-- Any topology coinduced by a locally path-connected topology is locally path-connected. -/ +lemma LocPathConnectedSpace.coinduced {Y : Type*} (f : X → Y) : + @LocPathConnectedSpace Y (.coinduced f ‹_›) := by + let _ := TopologicalSpace.coinduced f ‹_›; have hf : Continuous f := continuous_coinduced_rng + refine locPathConnectedSpace_iff_isOpen_pathComponentIn.mpr fun y u hu ↦ + isOpen_coinduced.mpr <| isOpen_iff_mem_nhds.mpr fun x hx ↦ ?_ + have hx' := preimage_mono pathComponentIn_subset hx + refine mem_nhds_iff.mpr ⟨pathComponentIn x (f ⁻¹' u), ?_, + (hu.preimage hf).pathComponentIn _, mem_pathComponentIn_self hx'⟩ + rw [← image_subset_iff, ← pathComponentIn_congr hx] + exact ((isPathConnected_pathComponentIn hx').image hf).subset_pathComponentIn + ⟨x, mem_pathComponentIn_self hx', rfl⟩ <| + (image_mono pathComponentIn_subset).trans <| u.image_preimage_subset f + +/-- Quotients of locally path-connected spaces are locally path-connected. -/ +lemma Topology.IsQuotientMap.locPathConnectedSpace {f : X → Y} (h : IsQuotientMap f) : + LocPathConnectedSpace Y := + h.2 ▸ LocPathConnectedSpace.coinduced f + +/-- Quotients of locally path-connected spaces are locally path-connected. -/ +instance Quot.locPathConnectedSpace {r : X → X → Prop} : LocPathConnectedSpace (Quot r) := + isQuotientMap_quot_mk.locPathConnectedSpace + +/-- Quotients of locally path-connected spaces are locally path-connected. -/ +instance Quotient.locPathConnectedSpace {s : Setoid X} : LocPathConnectedSpace (Quotient s) := + isQuotientMap_quotient_mk'.locPathConnectedSpace + + +/-- Disjoint unions of locally path-connected spaces are locally path-connected. -/ +instance Sum.locPathConnectedSpace.{u} {X Y : Type u} [TopologicalSpace X] [TopologicalSpace Y] + [LocPathConnectedSpace X] [LocPathConnectedSpace Y] : + LocPathConnectedSpace (X ⊕ Y) := by + rw [locPathConnectedSpace_iff_pathComponentIn_mem_nhds]; intro x u hu hxu; rw [mem_nhds_iff] + obtain x | y := x + · refine ⟨Sum.inl '' (pathComponentIn x (Sum.inl ⁻¹' u)), ?_, ?_, ?_⟩ + · apply IsPathConnected.subset_pathComponentIn + · exact (isPathConnected_pathComponentIn (by exact hxu)).image continuous_inl + · exact ⟨x, mem_pathComponentIn_self hxu, rfl⟩ + · exact (image_mono pathComponentIn_subset).trans (u.image_preimage_subset _) + · exact isOpenMap_inl _ <| (hu.preimage continuous_inl).pathComponentIn _ + · exact ⟨x, mem_pathComponentIn_self hxu, rfl⟩ + · refine ⟨Sum.inr '' (pathComponentIn y (Sum.inr ⁻¹' u)), ?_, ?_, ?_⟩ + · apply IsPathConnected.subset_pathComponentIn + · exact (isPathConnected_pathComponentIn (by exact hxu)).image continuous_inr + · exact ⟨y, mem_pathComponentIn_self hxu, rfl⟩ + · exact (image_mono pathComponentIn_subset).trans (u.image_preimage_subset _) + · exact isOpenMap_inr _ <| (hu.preimage continuous_inr).pathComponentIn _ + · exact ⟨y, mem_pathComponentIn_self hxu, rfl⟩ + + +/-- Disjoint unions of locally path-connected spaces are locally path-connected. -/ +instance Sigma.locPathConnectedSpace {X : ι → Type*} + [(i : ι) → TopologicalSpace (X i)] [(i : ι) → LocPathConnectedSpace (X i)] : + LocPathConnectedSpace ((i : ι) × X i) := by + rw [locPathConnectedSpace_iff_pathComponentIn_mem_nhds]; intro x u hu hxu; rw [mem_nhds_iff] + refine ⟨(Sigma.mk x.1) '' (pathComponentIn x.2 ((Sigma.mk x.1) ⁻¹' u)), ?_, ?_, ?_⟩ + · apply IsPathConnected.subset_pathComponentIn + · exact (isPathConnected_pathComponentIn (by exact hxu)).image continuous_sigmaMk + · exact ⟨x.2, mem_pathComponentIn_self hxu, rfl⟩ + · exact (image_mono pathComponentIn_subset).trans (u.image_preimage_subset _) + · exact isOpenMap_sigmaMk _ <| (hu.preimage continuous_sigmaMk).pathComponentIn _ + · exact ⟨x.2, mem_pathComponentIn_self hxu, rfl⟩ + end LocPathConnectedSpace From d6c2c9157d71b59d98033b31423a0db08f11c4b4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 19 Nov 2024 16:30:33 +0000 Subject: [PATCH 116/829] feat: inclusion-exclusion principle (#17957) It was a long-standing TODO to have this in mathlib. --- Mathlib.lean | 1 + .../Algebra/BigOperators/Group/Finset.lean | 6 + .../Enumerative/InclusionExclusion.lean | 129 ++++++++++++++++++ docs/100.yaml | 7 +- 4 files changed, 142 insertions(+), 1 deletion(-) create mode 100644 Mathlib/Combinatorics/Enumerative/InclusionExclusion.lean diff --git a/Mathlib.lean b/Mathlib.lean index 49dd52b773d65..3bde3993e0818 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2138,6 +2138,7 @@ import Mathlib.Combinatorics.Enumerative.Composition import Mathlib.Combinatorics.Enumerative.DoubleCounting import Mathlib.Combinatorics.Enumerative.DyckWord import Mathlib.Combinatorics.Enumerative.IncidenceAlgebra +import Mathlib.Combinatorics.Enumerative.InclusionExclusion import Mathlib.Combinatorics.Enumerative.Partition import Mathlib.Combinatorics.HalesJewett import Mathlib.Combinatorics.Hall.Basic diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index 29817d8370d97..a8d12379bbcce 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -420,6 +420,12 @@ theorem prod_filter_mul_prod_filter_not have := Classical.decEq α rw [← prod_union (disjoint_filter_filter_neg s s p), filter_union_filter_neg_eq] +@[to_additive] +lemma prod_filter_not_mul_prod_filter (s : Finset α) (p : α → Prop) [DecidablePred p] + [∀ x, Decidable (¬p x)] (f : α → β) : + (∏ x ∈ s.filter fun x ↦ ¬p x, f x) * ∏ x ∈ s.filter p, f x = ∏ x ∈ s, f x := by + rw [mul_comm, prod_filter_mul_prod_filter_not] + section ToList @[to_additive (attr := simp)] diff --git a/Mathlib/Combinatorics/Enumerative/InclusionExclusion.lean b/Mathlib/Combinatorics/Enumerative/InclusionExclusion.lean new file mode 100644 index 0000000000000..314e4eea70693 --- /dev/null +++ b/Mathlib/Combinatorics/Enumerative/InclusionExclusion.lean @@ -0,0 +1,129 @@ +/- +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.Pi +import Mathlib.Algebra.BigOperators.Ring +import Mathlib.Algebra.Module.BigOperators + +/-! +# Inclusion-exclusion principle + +This file proves several variants of the inclusion-exclusion principle. + +The inclusion-exclusion principle says that the sum/integral of a function over a finite union of +sets can be calculated as the alternating sum over `n > 0` of the sum/integral of the function over +the intersection of `n` of the sets. + +By taking complements, it also says that the sum/integral of a function over a finite intersection +of complements of sets can be calculated as the alternating sum over `n ≥ 0` of the sum/integral of +the function over the intersection of `n` of the sets. + +By taking the function to be constant `1`, we instead get a result about the cardinality/measure of +the sets. + +## Main declarations + +Per the above explanation, this file contains the following variants of inclusion-exclusion: +* `Finset.inclusion_exclusion_sum_biUnion`: Sum of a function over a finite union of sets +* `Finset.inclusion_exclusion_card_biUnion`: Cardinality of a finite union of sets +* `Finset.inclusion_exclusion_sum_inf_compl`: Sum of a function over a finite intersection of + complements of sets +* `Finset.inclusion_exclusion_card_inf_compl`: Cardinality of a finite intersection of + complements of sets + +## TODO + +* Use (a slight variant of) `Finset.prod_indicator_biUnion_sub_indicator` to prove the integral + version of inclusion-exclusion. +* Prove that truncating the series alternatively gives an upper/lower bound to the true value. +-/ + +namespace Finset +variable {ι α G : Type*} [DecidableEq α] [AddCommGroup G] {s : Finset ι} + +lemma prod_indicator_biUnion_sub_indicator (hs : s.Nonempty) (S : ι → Finset α) (a : α) : + ∏ i ∈ s, (Set.indicator (s.biUnion S) 1 a - Set.indicator (S i) 1 a) = (0 : ℤ) := by + by_cases ha : a ∈ s.biUnion S + · obtain ⟨i, hi, ha⟩ := mem_biUnion.1 ha + exact prod_eq_zero hi <| by simp [*, -coe_biUnion] + · obtain ⟨i, hi⟩ := hs + have ha : a ∉ S i := fun h ↦ ha <| subset_biUnion_of_mem _ hi h + exact prod_eq_zero hi <| by simp [*, -coe_biUnion] + +/-- **Inclusion-exclusion principle** for the sum of a function over a union. + +The sum of a function `f` over the union of the `S i` over `i ∈ s` is the alternating sum of the +sums of `f` over the intersections of the `S i`. -/ +theorem inclusion_exclusion_sum_biUnion (s : Finset ι) (S : ι → Finset α) (f : α → G) : + ∑ a ∈ s.biUnion S, f a = ∑ t : s.powerset.filter (·.Nonempty), + (-1) ^ (#t.1 + 1) • ∑ a ∈ t.1.inf' (mem_filter.1 t.2).2 S, f a := by + classical + rw [← sub_eq_zero] + calc + ∑ a ∈ s.biUnion S, f a - ∑ t : s.powerset.filter (·.Nonempty), + (-1) ^ (#t.1 + 1) • ∑ a ∈ t.1.inf' (mem_filter.1 t.2).2 S, f a + = ∑ t : s.powerset.filter (·.Nonempty), + (-1) ^ #t.1 • ∑ a ∈ t.1.inf' (mem_filter.1 t.2).2 S, f a + + ∑ t ∈ s.powerset.filter (¬ ·.Nonempty), (-1) ^ #t • ∑ a ∈ s.biUnion S, f a := by + simp [sub_eq_neg_add, ← sum_neg_distrib, filter_eq', pow_succ] + _ = ∑ t ∈ s.powerset, (-1) ^ #t • + if ht : t.Nonempty then ∑ a ∈ t.inf' ht S, f a else ∑ a ∈ s.biUnion S, f a := by + rw [← sum_attach (filter ..)]; simp [sum_dite, filter_eq', sum_attach] + _ = ∑ a ∈ s.biUnion S, (∏ i ∈ s, (1 - Set.indicator (S i) 1 a : ℤ)) • f a := by + simp only [Int.reduceNeg, mem_coe, prod_sub, sum_comm (s := s.biUnion S), sum_smul, mul_assoc] + congr! with t + split_ifs with ht + · obtain ⟨i, hi⟩ := ht + simp only [prod_const_one, mul_one, prod_indicator_apply] + simp only [smul_sum, Set.indicator, Set.mem_iInter, mem_coe, Pi.one_apply, mul_ite, mul_one, + mul_zero, ite_smul, zero_smul, sum_ite, not_forall, sum_const_zero, add_zero] + congr + aesop + · obtain rfl := not_nonempty_iff_eq_empty.1 ht + simp + _ = ∑ a ∈ s.biUnion S, (∏ i ∈ s, + (Set.indicator (s.biUnion S) 1 a - Set.indicator (S i) 1 a) : ℤ) • f a := by + congr! with t; rw [Set.indicator_of_mem ‹_›, Pi.one_apply] + _ = 0 := by + obtain rfl | hs := s.eq_empty_or_nonempty <;> + simp [-coe_biUnion, prod_indicator_biUnion_sub_indicator, *] + +/-- **Inclusion-exclusion principle** for the cardinality of a union. + +The cardinality of the union of the `S i` over `i ∈ s` is the alternating sum of the cardinalities +of the intersections of the `S i`. -/ +theorem inclusion_exclusion_card_biUnion (s : Finset ι) (S : ι → Finset α) : + #(s.biUnion S) = ∑ t : s.powerset.filter (·.Nonempty), + (-1 : ℤ) ^ (#t.1 + 1) * #(t.1.inf' (mem_filter.1 t.2).2 S) := by + simpa using inclusion_exclusion_sum_biUnion (G := ℤ) s S (f := 1) + +variable [Fintype α] + +/-- **Inclusion-exclusion principle** for the sum of a function over an intersection of complements. + +The sum of a function `f` over the intersection of the complements of the `S i` over `i ∈ s` is the +alternating sum of the sums of `f` over the intersections of the `S i`. -/ +theorem inclusion_exclusion_sum_inf_compl (s : Finset ι) (S : ι → Finset α) (f : α → G) : + ∑ a ∈ s.inf fun i ↦ (S i)ᶜ, f a = ∑ t ∈ s.powerset, (-1) ^ #t • ∑ a ∈ t.inf S, f a := by + classical + calc + ∑ a ∈ s.inf fun i ↦ (S i)ᶜ, f a + = ∑ a, f a - ∑ a ∈ s.biUnion S, f a := by + rw [← Finset.compl_sup, sup_eq_biUnion, eq_sub_iff_add_eq, sum_compl_add_sum] + _ = ∑ t ∈ s.powerset.filter (¬ ·.Nonempty), (-1) ^ #t • ∑ a ∈ t.inf S, f a + + ∑ t ∈ s.powerset.filter (·.Nonempty), (-1) ^ #t • ∑ a ∈ t.inf S, f a := by + simp [← sum_attach (filter ..), inclusion_exclusion_sum_biUnion, inf'_eq_inf, filter_eq', + sub_eq_add_neg, pow_succ] + _ = ∑ t ∈ s.powerset, (-1) ^ #t • ∑ a ∈ t.inf S, f a := sum_filter_not_add_sum_filter .. + +/-- **Inclusion-exclusion principle** for the cardinality of an intersection of complements. + +The cardinality of the intersection of the complements of the `S i` over `i ∈ s` is the +alternating sum of the cardinalities of the intersections of the `S i`. -/ +theorem inclusion_exclusion_card_inf_compl (s : Finset ι) (S : ι → Finset α) : + #(s.inf fun i ↦ (S i)ᶜ) = ∑ t ∈ s.powerset, (-1 : ℤ) ^ #t * #(t.inf S) := by + simpa using inclusion_exclusion_sum_inf_compl (G := ℤ) s S (f := 1) + +end Finset diff --git a/docs/100.yaml b/docs/100.yaml index 670ae3991e330..31f5837e872ca 100644 --- a/docs/100.yaml +++ b/docs/100.yaml @@ -395,9 +395,14 @@ author : Manuel Candales 96: title : Principle of Inclusion/Exclusion + decls : + - Finset.inclusion_exclusion_sum_biUnion + - Finset.inclusion_exclusion_card_biUnion + - Finset.inclusion_exclusion_sum_inf_compl + - Finset.inclusion_exclusion_card_inf_compl links : github : https://github.com/NeilStrickland/lean_lib/blob/f88d162da2f990b87c4d34f5f46bbca2bbc5948e/src/combinatorics/matching.lean#L304 - author : Neil Strickland + author : Neil Strickland (outside mathlib), Yaël Dillies (in mathlib) 97: title : Cramer’s Rule decl : Matrix.mulVec_cramer From 53523844a7ce22746461dc24cb05736dfe43bb91 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 19 Nov 2024 16:30:35 +0000 Subject: [PATCH 117/829] chore: split Topology.UniformSpace.Basic (#19194) and reorganize material in UniformSpace.Compact (moving some material from Basic into it, and splitting parts of the original content out into later files) --- Mathlib.lean | 2 + .../Topology/Algebra/UniformGroup/Basic.lean | 1 + .../Topology/Algebra/UniformGroup/Defs.lean | 2 +- Mathlib/Topology/EMetricSpace/Defs.lean | 1 + Mathlib/Topology/GDelta/UniformSpace.lean | 2 +- Mathlib/Topology/MetricSpace/Pseudo/Defs.lean | 1 + Mathlib/Topology/UniformSpace/Basic.lean | 99 +----- Mathlib/Topology/UniformSpace/Cauchy.lean | 2 +- Mathlib/Topology/UniformSpace/Compact.lean | 323 ++++++------------ .../UniformSpace/CompactConvergence.lean | 1 + .../Topology/UniformSpace/HeineCantor.lean | 146 ++++++++ .../Topology/UniformSpace/OfCompactT2.lean | 107 ++++++ Mathlib/Topology/UniformSpace/OfFun.lean | 2 +- Mathlib/Topology/UniformSpace/Separation.lean | 2 +- .../UniformSpace/UniformConvergence.lean | 1 - 15 files changed, 367 insertions(+), 325 deletions(-) create mode 100644 Mathlib/Topology/UniformSpace/HeineCantor.lean create mode 100644 Mathlib/Topology/UniformSpace/OfCompactT2.lean diff --git a/Mathlib.lean b/Mathlib.lean index 3bde3993e0818..866e1168d05e1 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -5245,7 +5245,9 @@ import Mathlib.Topology.UniformSpace.CompleteSeparated import Mathlib.Topology.UniformSpace.Completion import Mathlib.Topology.UniformSpace.Equicontinuity import Mathlib.Topology.UniformSpace.Equiv +import Mathlib.Topology.UniformSpace.HeineCantor import Mathlib.Topology.UniformSpace.Matrix +import Mathlib.Topology.UniformSpace.OfCompactT2 import Mathlib.Topology.UniformSpace.OfFun import Mathlib.Topology.UniformSpace.Pi import Mathlib.Topology.UniformSpace.Separation diff --git a/Mathlib/Topology/Algebra/UniformGroup/Basic.lean b/Mathlib/Topology/Algebra/UniformGroup/Basic.lean index 8fd7dd61a29d2..4baeea0bda780 100644 --- a/Mathlib/Topology/Algebra/UniformGroup/Basic.lean +++ b/Mathlib/Topology/Algebra/UniformGroup/Basic.lean @@ -7,6 +7,7 @@ import Mathlib.Topology.UniformSpace.UniformConvergence import Mathlib.Topology.UniformSpace.UniformEmbedding import Mathlib.Topology.UniformSpace.CompleteSeparated import Mathlib.Topology.UniformSpace.Compact +import Mathlib.Topology.UniformSpace.HeineCantor import Mathlib.Topology.Algebra.UniformGroup.Defs import Mathlib.Topology.Algebra.Group.Quotient import Mathlib.Topology.DiscreteSubset diff --git a/Mathlib/Topology/Algebra/UniformGroup/Defs.lean b/Mathlib/Topology/Algebra/UniformGroup/Defs.lean index b593b9d1eb3ca..3455d9f14d056 100644 --- a/Mathlib/Topology/Algebra/UniformGroup/Defs.lean +++ b/Mathlib/Topology/Algebra/UniformGroup/Defs.lean @@ -3,8 +3,8 @@ Copyright (c) 2018 Patrick Massot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Johannes Hölzl -/ -import Mathlib.Topology.UniformSpace.Basic import Mathlib.Topology.Algebra.Group.Basic +import Mathlib.Topology.UniformSpace.Basic /-! # Uniform structure on topological groups diff --git a/Mathlib/Topology/EMetricSpace/Defs.lean b/Mathlib/Topology/EMetricSpace/Defs.lean index 1c005ea711864..445f6803e42d1 100644 --- a/Mathlib/Topology/EMetricSpace/Defs.lean +++ b/Mathlib/Topology/EMetricSpace/Defs.lean @@ -5,6 +5,7 @@ Authors: Jeremy Avigad, Robert Y. Lewis, Johannes Hölzl, Mario Carneiro, Sébas -/ import Mathlib.Data.ENNReal.Inv import Mathlib.Topology.UniformSpace.OfFun +import Mathlib.Topology.Bases /-! # Extended metric spaces diff --git a/Mathlib/Topology/GDelta/UniformSpace.lean b/Mathlib/Topology/GDelta/UniformSpace.lean index 46bdc36f2fdab..3aab2f9dd00c0 100644 --- a/Mathlib/Topology/GDelta/UniformSpace.lean +++ b/Mathlib/Topology/GDelta/UniformSpace.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Yury Kudryashov -/ import Mathlib.Topology.GDelta.Basic -import Mathlib.Topology.UniformSpace.Basic import Mathlib.Order.Filter.CountableInter +import Mathlib.Topology.UniformSpace.Basic /-! # `Gδ` sets and uniform spaces diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean b/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean index 19689b2f5cdf6..82a80a1ff7f93 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean @@ -6,6 +6,7 @@ Authors: Jeremy Avigad, Robert Y. Lewis, Johannes Hölzl, Mario Carneiro, Sébas import Mathlib.Data.ENNReal.Real import Mathlib.Tactic.Bound.Attribute import Mathlib.Topology.EMetricSpace.Defs +import Mathlib.Topology.UniformSpace.Compact /-! ## Pseudo-metric spaces diff --git a/Mathlib/Topology/UniformSpace/Basic.lean b/Mathlib/Topology/UniformSpace/Basic.lean index 47f236596aa6b..5534b8ebfe86d 100644 --- a/Mathlib/Topology/UniformSpace/Basic.lean +++ b/Mathlib/Topology/UniformSpace/Basic.lean @@ -5,9 +5,10 @@ Authors: Johannes Hölzl, Mario Carneiro, Patrick Massot -/ import Mathlib.Order.Filter.SmallSets import Mathlib.Tactic.Monotonicity.Basic -import Mathlib.Topology.Compactness.Compact +import Mathlib.Order.Filter.Tendsto import Mathlib.Topology.NhdsSet import Mathlib.Algebra.Group.Defs +import Mathlib.Topology.ContinuousOn /-! # Uniform spaces @@ -1608,102 +1609,6 @@ end Sum end Constructions -/-! -### Compact sets in uniform spaces --/ - -section Compact - -open UniformSpace -variable [UniformSpace α] {K : Set α} - -/-- Let `c : ι → Set α` be an open cover of a compact set `s`. Then there exists an entourage -`n` such that for each `x ∈ s` its `n`-neighborhood is contained in some `c i`. -/ -theorem lebesgue_number_lemma {ι : Sort*} {U : ι → Set α} (hK : IsCompact K) - (hopen : ∀ i, IsOpen (U i)) (hcover : K ⊆ ⋃ i, U i) : - ∃ V ∈ 𝓤 α, ∀ x ∈ K, ∃ i, ball x V ⊆ U i := by - have : ∀ x ∈ K, ∃ i, ∃ V ∈ 𝓤 α, ball x (V ○ V) ⊆ U i := fun x hx ↦ by - obtain ⟨i, hi⟩ := mem_iUnion.1 (hcover hx) - rw [← (hopen i).mem_nhds_iff, nhds_eq_comap_uniformity, ← lift'_comp_uniformity] at hi - exact ⟨i, (((basis_sets _).lift' <| monotone_id.compRel monotone_id).comap _).mem_iff.1 hi⟩ - choose ind W hW hWU using this - rcases hK.elim_nhds_subcover' (fun x hx ↦ ball x (W x hx)) (fun x hx ↦ ball_mem_nhds _ (hW x hx)) - with ⟨t, ht⟩ - refine ⟨⋂ x ∈ t, W x x.2, (biInter_finset_mem _).2 fun x _ ↦ hW x x.2, fun x hx ↦ ?_⟩ - rcases mem_iUnion₂.1 (ht hx) with ⟨y, hyt, hxy⟩ - exact ⟨ind y y.2, fun z hz ↦ hWU _ _ ⟨x, hxy, mem_iInter₂.1 hz _ hyt⟩⟩ - -/-- Let `U : ι → Set α` be an open cover of a compact set `K`. -Then there exists an entourage `V` -such that for each `x ∈ K` its `V`-neighborhood is included in some `U i`. - -Moreover, one can choose an entourage from a given basis. -/ -protected theorem Filter.HasBasis.lebesgue_number_lemma {ι' ι : Sort*} {p : ι' → Prop} - {V : ι' → Set (α × α)} {U : ι → Set α} (hbasis : (𝓤 α).HasBasis p V) (hK : IsCompact K) - (hopen : ∀ j, IsOpen (U j)) (hcover : K ⊆ ⋃ j, U j) : - ∃ i, p i ∧ ∀ x ∈ K, ∃ j, ball x (V i) ⊆ U j := by - refine (hbasis.exists_iff ?_).1 (lebesgue_number_lemma hK hopen hcover) - exact fun s t hst ht x hx ↦ (ht x hx).imp fun i hi ↦ Subset.trans (ball_mono hst _) hi - -/-- Let `c : Set (Set α)` be an open cover of a compact set `s`. Then there exists an entourage -`n` such that for each `x ∈ s` its `n`-neighborhood is contained in some `t ∈ c`. -/ -theorem lebesgue_number_lemma_sUnion {S : Set (Set α)} - (hK : IsCompact K) (hopen : ∀ s ∈ S, IsOpen s) (hcover : K ⊆ ⋃₀ S) : - ∃ V ∈ 𝓤 α, ∀ x ∈ K, ∃ s ∈ S, ball x V ⊆ s := by - rw [sUnion_eq_iUnion] at hcover - simpa using lebesgue_number_lemma hK (by simpa) hcover - -/-- If `K` is a compact set in a uniform space and `{V i | p i}` is a basis of entourages, -then `{⋃ x ∈ K, UniformSpace.ball x (V i) | p i}` is a basis of `𝓝ˢ K`. - -Here "`{s i | p i}` is a basis of a filter `l`" means `Filter.HasBasis l p s`. -/ -theorem IsCompact.nhdsSet_basis_uniformity {p : ι → Prop} {V : ι → Set (α × α)} - (hbasis : (𝓤 α).HasBasis p V) (hK : IsCompact K) : - (𝓝ˢ K).HasBasis p fun i => ⋃ x ∈ K, ball x (V i) where - mem_iff' U := by - constructor - · intro H - have HKU : K ⊆ ⋃ _ : Unit, interior U := by - simpa only [iUnion_const, subset_interior_iff_mem_nhdsSet] using H - obtain ⟨i, hpi, hi⟩ : ∃ i, p i ∧ ⋃ x ∈ K, ball x (V i) ⊆ interior U := by - simpa using hbasis.lebesgue_number_lemma hK (fun _ ↦ isOpen_interior) HKU - exact ⟨i, hpi, hi.trans interior_subset⟩ - · rintro ⟨i, hpi, hi⟩ - refine mem_of_superset (bUnion_mem_nhdsSet fun x _ ↦ ?_) hi - exact ball_mem_nhds _ <| hbasis.mem_of_mem hpi - --- TODO: move to a separate file, golf using the regularity of a uniform space. -theorem Disjoint.exists_uniform_thickening {A B : Set α} (hA : IsCompact A) (hB : IsClosed B) - (h : Disjoint A B) : ∃ V ∈ 𝓤 α, Disjoint (⋃ x ∈ A, ball x V) (⋃ x ∈ B, ball x V) := by - have : Bᶜ ∈ 𝓝ˢ A := hB.isOpen_compl.mem_nhdsSet.mpr h.le_compl_right - rw [(hA.nhdsSet_basis_uniformity (Filter.basis_sets _)).mem_iff] at this - rcases this with ⟨U, hU, hUAB⟩ - rcases comp_symm_mem_uniformity_sets hU with ⟨V, hV, hVsymm, hVU⟩ - refine ⟨V, hV, Set.disjoint_left.mpr fun x => ?_⟩ - simp only [mem_iUnion₂] - rintro ⟨a, ha, hxa⟩ ⟨b, hb, hxb⟩ - rw [mem_ball_symmetry hVsymm] at hxa hxb - exact hUAB (mem_iUnion₂_of_mem ha <| hVU <| mem_comp_of_mem_ball hVsymm hxa hxb) hb - -theorem Disjoint.exists_uniform_thickening_of_basis {p : ι → Prop} {s : ι → Set (α × α)} - (hU : (𝓤 α).HasBasis p s) {A B : Set α} (hA : IsCompact A) (hB : IsClosed B) - (h : Disjoint A B) : ∃ i, p i ∧ Disjoint (⋃ x ∈ A, ball x (s i)) (⋃ x ∈ B, ball x (s i)) := by - rcases h.exists_uniform_thickening hA hB with ⟨V, hV, hVAB⟩ - rcases hU.mem_iff.1 hV with ⟨i, hi, hiV⟩ - exact ⟨i, hi, hVAB.mono (iUnion₂_mono fun a _ => ball_mono hiV a) - (iUnion₂_mono fun b _ => ball_mono hiV b)⟩ - -/-- A useful consequence of the Lebesgue number lemma: given any compact set `K` contained in an -open set `U`, we can find an (open) entourage `V` such that the ball of size `V` about any point of -`K` is contained in `U`. -/ -theorem lebesgue_number_of_compact_open {K U : Set α} (hK : IsCompact K) - (hU : IsOpen U) (hKU : K ⊆ U) : ∃ V ∈ 𝓤 α, IsOpen V ∧ ∀ x ∈ K, UniformSpace.ball x V ⊆ U := - let ⟨V, ⟨hV, hVo⟩, hVU⟩ := - (hK.nhdsSet_basis_uniformity uniformity_hasBasis_open).mem_iff.1 (hU.mem_nhdsSet.2 hKU) - ⟨V, hV, hVo, iUnion₂_subset_iff.1 hVU⟩ - -end Compact - /-! ### Expressing continuity properties in uniform spaces diff --git a/Mathlib/Topology/UniformSpace/Cauchy.lean b/Mathlib/Topology/UniformSpace/Cauchy.lean index 9962bbe2a00d6..b8ae4225ddef4 100644 --- a/Mathlib/Topology/UniformSpace/Cauchy.lean +++ b/Mathlib/Topology/UniformSpace/Cauchy.lean @@ -5,8 +5,8 @@ Authors: Johannes Hölzl, Mario Carneiro -/ import Mathlib.Topology.Algebra.Constructions import Mathlib.Topology.Bases -import Mathlib.Topology.UniformSpace.Basic import Mathlib.Algebra.Order.Group.Nat +import Mathlib.Topology.UniformSpace.Basic /-! # Theory of Cauchy filters in uniform spaces. Complete uniform spaces. Totally bounded subsets. diff --git a/Mathlib/Topology/UniformSpace/Compact.lean b/Mathlib/Topology/UniformSpace/Compact.lean index f8308127388e0..ed4c24be25ac2 100644 --- a/Mathlib/Topology/UniformSpace/Compact.lean +++ b/Mathlib/Topology/UniformSpace/Compact.lean @@ -1,44 +1,115 @@ /- -Copyright (c) 2020 Patrick Massot. All rights reserved. +Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Patrick Massot, Yury Kudryashov +Authors: Johannes Hölzl, Mario Carneiro, Patrick Massot -/ -import Mathlib.Topology.UniformSpace.UniformConvergence -import Mathlib.Topology.UniformSpace.Equicontinuity -import Mathlib.Topology.Support +import Mathlib.Topology.UniformSpace.Basic +import Mathlib.Topology.Compactness.Compact /-! -# Compact separated uniform spaces - -## Main statements +# Compact sets in uniform spaces * `compactSpace_uniformity`: On a compact uniform space, the topology determines the uniform structure, entourages are exactly the neighborhoods of the diagonal. -* `uniformSpace_of_compact_t2`: every compact T2 topological structure is induced by a uniform - structure. This uniform structure is described in the previous item. - -* **Heine-Cantor** theorem: continuous functions on compact uniform spaces with values in uniform - spaces are automatically uniformly continuous. There are several variations, the main one is - `CompactSpace.uniformContinuous_of_continuous`. - -## Implementation notes - -The construction `uniformSpace_of_compact_t2` is not declared as an instance, as it would badly -loop. - -## Tags - -uniform space, uniform continuity, compact space -/ -open Uniformity Topology Filter UniformSpace Set +universe u v ua ub uc ud + +variable {α : Type ua} {β : Type ub} {γ : Type uc} {δ : Type ud} {ι : Sort*} + +section Compact + +open Uniformity Set Filter UniformSpace +open scoped Topology + +variable [UniformSpace α] {K : Set α} + +/-- Let `c : ι → Set α` be an open cover of a compact set `s`. Then there exists an entourage +`n` such that for each `x ∈ s` its `n`-neighborhood is contained in some `c i`. -/ +theorem lebesgue_number_lemma {ι : Sort*} {U : ι → Set α} (hK : IsCompact K) + (hopen : ∀ i, IsOpen (U i)) (hcover : K ⊆ ⋃ i, U i) : + ∃ V ∈ 𝓤 α, ∀ x ∈ K, ∃ i, ball x V ⊆ U i := by + have : ∀ x ∈ K, ∃ i, ∃ V ∈ 𝓤 α, ball x (V ○ V) ⊆ U i := fun x hx ↦ by + obtain ⟨i, hi⟩ := mem_iUnion.1 (hcover hx) + rw [← (hopen i).mem_nhds_iff, nhds_eq_comap_uniformity, ← lift'_comp_uniformity] at hi + exact ⟨i, (((basis_sets _).lift' <| monotone_id.compRel monotone_id).comap _).mem_iff.1 hi⟩ + choose ind W hW hWU using this + rcases hK.elim_nhds_subcover' (fun x hx ↦ ball x (W x hx)) (fun x hx ↦ ball_mem_nhds _ (hW x hx)) + with ⟨t, ht⟩ + refine ⟨⋂ x ∈ t, W x x.2, (biInter_finset_mem _).2 fun x _ ↦ hW x x.2, fun x hx ↦ ?_⟩ + rcases mem_iUnion₂.1 (ht hx) with ⟨y, hyt, hxy⟩ + exact ⟨ind y y.2, fun z hz ↦ hWU _ _ ⟨x, hxy, mem_iInter₂.1 hz _ hyt⟩⟩ + +/-- Let `U : ι → Set α` be an open cover of a compact set `K`. +Then there exists an entourage `V` +such that for each `x ∈ K` its `V`-neighborhood is included in some `U i`. + +Moreover, one can choose an entourage from a given basis. -/ +protected theorem Filter.HasBasis.lebesgue_number_lemma {ι' ι : Sort*} {p : ι' → Prop} + {V : ι' → Set (α × α)} {U : ι → Set α} (hbasis : (𝓤 α).HasBasis p V) (hK : IsCompact K) + (hopen : ∀ j, IsOpen (U j)) (hcover : K ⊆ ⋃ j, U j) : + ∃ i, p i ∧ ∀ x ∈ K, ∃ j, ball x (V i) ⊆ U j := by + refine (hbasis.exists_iff ?_).1 (lebesgue_number_lemma hK hopen hcover) + exact fun s t hst ht x hx ↦ (ht x hx).imp fun i hi ↦ Subset.trans (ball_mono hst _) hi + +/-- Let `c : Set (Set α)` be an open cover of a compact set `s`. Then there exists an entourage +`n` such that for each `x ∈ s` its `n`-neighborhood is contained in some `t ∈ c`. -/ +theorem lebesgue_number_lemma_sUnion {S : Set (Set α)} + (hK : IsCompact K) (hopen : ∀ s ∈ S, IsOpen s) (hcover : K ⊆ ⋃₀ S) : + ∃ V ∈ 𝓤 α, ∀ x ∈ K, ∃ s ∈ S, ball x V ⊆ s := by + rw [sUnion_eq_iUnion] at hcover + simpa using lebesgue_number_lemma hK (by simpa) hcover + +/-- If `K` is a compact set in a uniform space and `{V i | p i}` is a basis of entourages, +then `{⋃ x ∈ K, UniformSpace.ball x (V i) | p i}` is a basis of `𝓝ˢ K`. + +Here "`{s i | p i}` is a basis of a filter `l`" means `Filter.HasBasis l p s`. -/ +theorem IsCompact.nhdsSet_basis_uniformity {p : ι → Prop} {V : ι → Set (α × α)} + (hbasis : (𝓤 α).HasBasis p V) (hK : IsCompact K) : + (𝓝ˢ K).HasBasis p fun i => ⋃ x ∈ K, ball x (V i) where + mem_iff' U := by + constructor + · intro H + have HKU : K ⊆ ⋃ _ : Unit, interior U := by + simpa only [iUnion_const, subset_interior_iff_mem_nhdsSet] using H + obtain ⟨i, hpi, hi⟩ : ∃ i, p i ∧ ⋃ x ∈ K, ball x (V i) ⊆ interior U := by + simpa using hbasis.lebesgue_number_lemma hK (fun _ ↦ isOpen_interior) HKU + exact ⟨i, hpi, hi.trans interior_subset⟩ + · rintro ⟨i, hpi, hi⟩ + refine mem_of_superset (bUnion_mem_nhdsSet fun x _ ↦ ?_) hi + exact ball_mem_nhds _ <| hbasis.mem_of_mem hpi + +-- TODO: move to a separate file, golf using the regularity of a uniform space. +theorem Disjoint.exists_uniform_thickening {A B : Set α} (hA : IsCompact A) (hB : IsClosed B) + (h : Disjoint A B) : ∃ V ∈ 𝓤 α, Disjoint (⋃ x ∈ A, ball x V) (⋃ x ∈ B, ball x V) := by + have : Bᶜ ∈ 𝓝ˢ A := hB.isOpen_compl.mem_nhdsSet.mpr h.le_compl_right + rw [(hA.nhdsSet_basis_uniformity (Filter.basis_sets _)).mem_iff] at this + rcases this with ⟨U, hU, hUAB⟩ + rcases comp_symm_mem_uniformity_sets hU with ⟨V, hV, hVsymm, hVU⟩ + refine ⟨V, hV, Set.disjoint_left.mpr fun x => ?_⟩ + simp only [mem_iUnion₂] + rintro ⟨a, ha, hxa⟩ ⟨b, hb, hxb⟩ + rw [mem_ball_symmetry hVsymm] at hxa hxb + exact hUAB (mem_iUnion₂_of_mem ha <| hVU <| mem_comp_of_mem_ball hVsymm hxa hxb) hb + +theorem Disjoint.exists_uniform_thickening_of_basis {p : ι → Prop} {s : ι → Set (α × α)} + (hU : (𝓤 α).HasBasis p s) {A B : Set α} (hA : IsCompact A) (hB : IsClosed B) + (h : Disjoint A B) : ∃ i, p i ∧ Disjoint (⋃ x ∈ A, ball x (s i)) (⋃ x ∈ B, ball x (s i)) := by + rcases h.exists_uniform_thickening hA hB with ⟨V, hV, hVAB⟩ + rcases hU.mem_iff.1 hV with ⟨i, hi, hiV⟩ + exact ⟨i, hi, hVAB.mono (iUnion₂_mono fun a _ => ball_mono hiV a) + (iUnion₂_mono fun b _ => ball_mono hiV b)⟩ + +/-- A useful consequence of the Lebesgue number lemma: given any compact set `K` contained in an +open set `U`, we can find an (open) entourage `V` such that the ball of size `V` about any point of +`K` is contained in `U`. -/ +theorem lebesgue_number_of_compact_open {K U : Set α} (hK : IsCompact K) + (hU : IsOpen U) (hKU : K ⊆ U) : ∃ V ∈ 𝓤 α, IsOpen V ∧ ∀ x ∈ K, UniformSpace.ball x V ⊆ U := + let ⟨V, ⟨hV, hVo⟩, hVU⟩ := + (hK.nhdsSet_basis_uniformity uniformity_hasBasis_open).mem_iff.1 (hU.mem_nhdsSet.2 hKU) + ⟨V, hV, hVo, iUnion₂_subset_iff.1 hVU⟩ -variable {α β γ : Type*} [UniformSpace α] [UniformSpace β] - -/-! -### Uniformity on compact spaces --/ /-- On a compact uniform space, the topology determines the uniform structure, entourages are exactly the neighborhoods of the diagonal. -/ @@ -66,196 +137,4 @@ theorem unique_uniformity_of_compact [t : TopologicalSpace γ] [CompactSpace γ] have : @CompactSpace γ u'.toTopologicalSpace := by rwa [h'] rw [@compactSpace_uniformity _ u, compactSpace_uniformity, h, h'] -/-- The unique uniform structure inducing a given compact topological structure. -/ -def uniformSpaceOfCompactT2 [TopologicalSpace γ] [CompactSpace γ] [T2Space γ] : UniformSpace γ where - uniformity := 𝓝ˢ (diagonal γ) - symm := continuous_swap.tendsto_nhdsSet fun _ => Eq.symm - comp := by - /- This is the difficult part of the proof. We need to prove that, for each neighborhood `W` - of the diagonal `Δ`, there exists a smaller neighborhood `V` such that `V ○ V ⊆ W`. - -/ - set 𝓝Δ := 𝓝ˢ (diagonal γ) - -- The filter of neighborhoods of Δ - set F := 𝓝Δ.lift' fun s : Set (γ × γ) => s ○ s - -- Compositions of neighborhoods of Δ - -- If this weren't true, then there would be V ∈ 𝓝Δ such that F ⊓ 𝓟 Vᶜ ≠ ⊥ - rw [le_iff_forall_inf_principal_compl] - intro V V_in - by_contra H - haveI : NeBot (F ⊓ 𝓟 Vᶜ) := ⟨H⟩ - -- Hence compactness would give us a cluster point (x, y) for F ⊓ 𝓟 Vᶜ - obtain ⟨⟨x, y⟩, hxy⟩ : ∃ p : γ × γ, ClusterPt p (F ⊓ 𝓟 Vᶜ) := exists_clusterPt_of_compactSpace _ - -- In particular (x, y) is a cluster point of 𝓟 Vᶜ, hence is not in the interior of V, - -- and a fortiori not in Δ, so x ≠ y - have clV : ClusterPt (x, y) (𝓟 <| Vᶜ) := hxy.of_inf_right - have : (x, y) ∉ interior V := by - have : (x, y) ∈ closure Vᶜ := by rwa [mem_closure_iff_clusterPt] - rwa [closure_compl] at this - have diag_subset : diagonal γ ⊆ interior V := subset_interior_iff_mem_nhdsSet.2 V_in - have x_ne_y : x ≠ y := mt (@diag_subset (x, y)) this - -- Since γ is compact and Hausdorff, it is T₄, hence T₃. - -- So there are closed neighborhoods V₁ and V₂ of x and y contained in - -- disjoint open neighborhoods U₁ and U₂. - obtain - ⟨U₁, _, V₁, V₁_in, U₂, _, V₂, V₂_in, V₁_cl, V₂_cl, U₁_op, U₂_op, VU₁, VU₂, hU₁₂⟩ := - disjoint_nested_nhds x_ne_y - -- We set U₃ := (V₁ ∪ V₂)ᶜ so that W := U₁ ×ˢ U₁ ∪ U₂ ×ˢ U₂ ∪ U₃ ×ˢ U₃ is an open - -- neighborhood of Δ. - let U₃ := (V₁ ∪ V₂)ᶜ - have U₃_op : IsOpen U₃ := (V₁_cl.union V₂_cl).isOpen_compl - let W := U₁ ×ˢ U₁ ∪ U₂ ×ˢ U₂ ∪ U₃ ×ˢ U₃ - have W_in : W ∈ 𝓝Δ := by - rw [mem_nhdsSet_iff_forall] - rintro ⟨z, z'⟩ (rfl : z = z') - refine IsOpen.mem_nhds ?_ ?_ - · apply_rules [IsOpen.union, IsOpen.prod] - · simp only [W, mem_union, mem_prod, and_self_iff] - exact (_root_.em _).imp_left fun h => union_subset_union VU₁ VU₂ h - -- So W ○ W ∈ F by definition of F - have : W ○ W ∈ F := @mem_lift' _ _ _ (fun s => s ○ s) _ W_in - -- Porting note: was `by simpa only using mem_lift' W_in` - -- And V₁ ×ˢ V₂ ∈ 𝓝 (x, y) - have hV₁₂ : V₁ ×ˢ V₂ ∈ 𝓝 (x, y) := prod_mem_nhds V₁_in V₂_in - -- But (x, y) is also a cluster point of F so (V₁ ×ˢ V₂) ∩ (W ○ W) ≠ ∅ - -- However the construction of W implies (V₁ ×ˢ V₂) ∩ (W ○ W) = ∅. - -- Indeed assume for contradiction there is some (u, v) in the intersection. - obtain ⟨⟨u, v⟩, ⟨u_in, v_in⟩, w, huw, hwv⟩ := clusterPt_iff.mp hxy.of_inf_left hV₁₂ this - -- So u ∈ V₁, v ∈ V₂, and there exists some w such that (u, w) ∈ W and (w ,v) ∈ W. - -- Because u is in V₁ which is disjoint from U₂ and U₃, (u, w) ∈ W forces (u, w) ∈ U₁ ×ˢ U₁. - have uw_in : (u, w) ∈ U₁ ×ˢ U₁ := - (huw.resolve_right fun h => h.1 <| Or.inl u_in).resolve_right fun h => - hU₁₂.le_bot ⟨VU₁ u_in, h.1⟩ - -- Similarly, because v ∈ V₂, (w ,v) ∈ W forces (w, v) ∈ U₂ ×ˢ U₂. - have wv_in : (w, v) ∈ U₂ ×ˢ U₂ := - (hwv.resolve_right fun h => h.2 <| Or.inr v_in).resolve_left fun h => - hU₁₂.le_bot ⟨h.2, VU₂ v_in⟩ - -- Hence w ∈ U₁ ∩ U₂ which is empty. - -- So we have a contradiction - exact hU₁₂.le_bot ⟨uw_in.2, wv_in.1⟩ - nhds_eq_comap_uniformity x := by - simp_rw [nhdsSet_diagonal, comap_iSup, nhds_prod_eq, comap_prod, Function.comp_def, comap_id'] - rw [iSup_split_single _ x, comap_const_of_mem fun V => mem_of_mem_nhds] - suffices ∀ y ≠ x, comap (fun _ : γ ↦ x) (𝓝 y) ⊓ 𝓝 y ≤ 𝓝 x by simpa - intro y hxy - simp [comap_const_of_not_mem (compl_singleton_mem_nhds hxy) (not_not_intro rfl)] - -/-! -### Heine-Cantor theorem --/ - - -/-- Heine-Cantor: a continuous function on a compact uniform space is uniformly -continuous. -/ -theorem CompactSpace.uniformContinuous_of_continuous [CompactSpace α] {f : α → β} - (h : Continuous f) : UniformContinuous f := -calc map (Prod.map f f) (𝓤 α) - = map (Prod.map f f) (𝓝ˢ (diagonal α)) := by rw [nhdsSet_diagonal_eq_uniformity] - _ ≤ 𝓝ˢ (diagonal β) := (h.prodMap h).tendsto_nhdsSet mapsTo_prod_map_diagonal - _ ≤ 𝓤 β := nhdsSet_diagonal_le_uniformity - -/-- Heine-Cantor: a continuous function on a compact set of a uniform space is uniformly -continuous. -/ -theorem IsCompact.uniformContinuousOn_of_continuous {s : Set α} {f : α → β} (hs : IsCompact s) - (hf : ContinuousOn f s) : UniformContinuousOn f s := by - rw [uniformContinuousOn_iff_restrict] - rw [isCompact_iff_compactSpace] at hs - rw [continuousOn_iff_continuous_restrict] at hf - exact CompactSpace.uniformContinuous_of_continuous hf - -/-- If `s` is compact and `f` is continuous at all points of `s`, then `f` is -"uniformly continuous at the set `s`", i.e. `f x` is close to `f y` whenever `x ∈ s` and `y` is -close to `x` (even if `y` is not itself in `s`, so this is a stronger assertion than -`UniformContinuousOn s`). -/ -theorem IsCompact.uniformContinuousAt_of_continuousAt {r : Set (β × β)} {s : Set α} - (hs : IsCompact s) (f : α → β) (hf : ∀ a ∈ s, ContinuousAt f a) (hr : r ∈ 𝓤 β) : - { x : α × α | x.1 ∈ s → (f x.1, f x.2) ∈ r } ∈ 𝓤 α := by - obtain ⟨t, ht, htsymm, htr⟩ := comp_symm_mem_uniformity_sets hr - choose U hU T hT hb using fun a ha => - exists_mem_nhds_ball_subset_of_mem_nhds ((hf a ha).preimage_mem_nhds <| mem_nhds_left _ ht) - obtain ⟨fs, hsU⟩ := hs.elim_nhds_subcover' U hU - apply mem_of_superset ((biInter_finset_mem fs).2 fun a _ => hT a a.2) - rintro ⟨a₁, a₂⟩ h h₁ - obtain ⟨a, ha, haU⟩ := Set.mem_iUnion₂.1 (hsU h₁) - apply htr - refine ⟨f a, htsymm.mk_mem_comm.1 (hb _ _ _ haU ?_), hb _ _ _ haU ?_⟩ - exacts [mem_ball_self _ (hT a a.2), mem_iInter₂.1 h a ha] - -theorem Continuous.uniformContinuous_of_tendsto_cocompact {f : α → β} {x : β} - (h_cont : Continuous f) (hx : Tendsto f (cocompact α) (𝓝 x)) : UniformContinuous f := - uniformContinuous_def.2 fun r hr => by - obtain ⟨t, ht, htsymm, htr⟩ := comp_symm_mem_uniformity_sets hr - obtain ⟨s, hs, hst⟩ := mem_cocompact.1 (hx <| mem_nhds_left _ ht) - apply - mem_of_superset - (symmetrize_mem_uniformity <| - (hs.uniformContinuousAt_of_continuousAt f fun _ _ => h_cont.continuousAt) <| - symmetrize_mem_uniformity hr) - rintro ⟨b₁, b₂⟩ h - by_cases h₁ : b₁ ∈ s; · exact (h.1 h₁).1 - by_cases h₂ : b₂ ∈ s; · exact (h.2 h₂).2 - apply htr - exact ⟨x, htsymm.mk_mem_comm.1 (hst h₁), hst h₂⟩ - -@[to_additive] -theorem HasCompactMulSupport.uniformContinuous_of_continuous {f : α → β} [One β] - (h1 : HasCompactMulSupport f) (h2 : Continuous f) : UniformContinuous f := - h2.uniformContinuous_of_tendsto_cocompact h1.is_one_at_infty - -/-- A family of functions `α → β → γ` tends uniformly to its value at `x` if `α` is locally compact, -`β` is compact and `f` is continuous on `U × (univ : Set β)` for some neighborhood `U` of `x`. -/ -theorem ContinuousOn.tendstoUniformly [LocallyCompactSpace α] [CompactSpace β] [UniformSpace γ] - {f : α → β → γ} {x : α} {U : Set α} (hxU : U ∈ 𝓝 x) (h : ContinuousOn (↿f) (U ×ˢ univ)) : - TendstoUniformly f (f x) (𝓝 x) := by - rcases LocallyCompactSpace.local_compact_nhds _ _ hxU with ⟨K, hxK, hKU, hK⟩ - have : UniformContinuousOn (↿f) (K ×ˢ univ) := - IsCompact.uniformContinuousOn_of_continuous (hK.prod isCompact_univ) - (h.mono <| prod_mono hKU Subset.rfl) - exact this.tendstoUniformly hxK - -/-- A continuous family of functions `α → β → γ` tends uniformly to its value at `x` -if `α` is weakly locally compact and `β` is compact. -/ -theorem Continuous.tendstoUniformly [WeaklyLocallyCompactSpace α] [CompactSpace β] [UniformSpace γ] - (f : α → β → γ) (h : Continuous ↿f) (x : α) : TendstoUniformly f (f x) (𝓝 x) := - let ⟨K, hK, hxK⟩ := exists_compact_mem_nhds x - have : UniformContinuousOn (↿f) (K ×ˢ univ) := - IsCompact.uniformContinuousOn_of_continuous (hK.prod isCompact_univ) h.continuousOn - this.tendstoUniformly hxK - -/-- In a product space `α × β`, assume that a function `f` is continuous on `s × k` where `k` is -compact. Then, along the fiber above any `q ∈ s`, `f` is transversely uniformly continuous, i.e., -if `p ∈ s` is close enough to `q`, then `f p x` is uniformly close to `f q x` for all `x ∈ k`. -/ -lemma IsCompact.mem_uniformity_of_prod - {α β E : Type*} [TopologicalSpace α] [TopologicalSpace β] [UniformSpace E] - {f : α → β → E} {s : Set α} {k : Set β} {q : α} {u : Set (E × E)} - (hk : IsCompact k) (hf : ContinuousOn f.uncurry (s ×ˢ k)) (hq : q ∈ s) (hu : u ∈ 𝓤 E) : - ∃ v ∈ 𝓝[s] q, ∀ p ∈ v, ∀ x ∈ k, (f p x, f q x) ∈ u := by - apply hk.induction_on (p := fun t ↦ ∃ v ∈ 𝓝[s] q, ∀ p ∈ v, ∀ x ∈ t, (f p x, f q x) ∈ u) - · exact ⟨univ, univ_mem, by simp⟩ - · intro t' t ht't ⟨v, v_mem, hv⟩ - exact ⟨v, v_mem, fun p hp x hx ↦ hv p hp x (ht't hx)⟩ - · intro t t' ⟨v, v_mem, hv⟩ ⟨v', v'_mem, hv'⟩ - refine ⟨v ∩ v', inter_mem v_mem v'_mem, fun p hp x hx ↦ ?_⟩ - rcases hx with h'x|h'x - · exact hv p hp.1 x h'x - · exact hv' p hp.2 x h'x - · rcases comp_symm_of_uniformity hu with ⟨u', u'_mem, u'_symm, hu'⟩ - intro x hx - obtain ⟨v, hv, w, hw, hvw⟩ : - ∃ v ∈ 𝓝[s] q, ∃ w ∈ 𝓝[k] x, v ×ˢ w ⊆ f.uncurry ⁻¹' {z | (f q x, z) ∈ u'} := - mem_nhdsWithin_prod_iff.1 (hf (q, x) ⟨hq, hx⟩ (mem_nhds_left (f q x) u'_mem)) - refine ⟨w, hw, v, hv, fun p hp y hy ↦ ?_⟩ - have A : (f q x, f p y) ∈ u' := hvw (⟨hp, hy⟩ : (p, y) ∈ v ×ˢ w) - have B : (f q x, f q y) ∈ u' := hvw (⟨mem_of_mem_nhdsWithin hq hv, hy⟩ : (q, y) ∈ v ×ˢ w) - exact hu' (prod_mk_mem_compRel (u'_symm A) B) - -section UniformConvergence - -/-- An equicontinuous family of functions defined on a compact uniform space is automatically -uniformly equicontinuous. -/ -theorem CompactSpace.uniformEquicontinuous_of_equicontinuous {ι : Type*} {F : ι → β → α} - [CompactSpace β] (h : Equicontinuous F) : UniformEquicontinuous F := by - rw [equicontinuous_iff_continuous] at h - rw [uniformEquicontinuous_iff_uniformContinuous] - exact CompactSpace.uniformContinuous_of_continuous h - -end UniformConvergence +end Compact diff --git a/Mathlib/Topology/UniformSpace/CompactConvergence.lean b/Mathlib/Topology/UniformSpace/CompactConvergence.lean index e39e65d2fbda4..d5e4df759ebef 100644 --- a/Mathlib/Topology/UniformSpace/CompactConvergence.lean +++ b/Mathlib/Topology/UniformSpace/CompactConvergence.lean @@ -7,6 +7,7 @@ import Mathlib.Topology.CompactOpen import Mathlib.Topology.LocallyFinite import Mathlib.Topology.Maps.Proper.Basic import Mathlib.Topology.UniformSpace.UniformConvergenceTopology +import Mathlib.Topology.UniformSpace.Compact /-! # Compact convergence (uniform convergence on compact sets) diff --git a/Mathlib/Topology/UniformSpace/HeineCantor.lean b/Mathlib/Topology/UniformSpace/HeineCantor.lean new file mode 100644 index 0000000000000..d97593eba0566 --- /dev/null +++ b/Mathlib/Topology/UniformSpace/HeineCantor.lean @@ -0,0 +1,146 @@ +/- +Copyright (c) 2020 Patrick Massot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Patrick Massot, Yury Kudryashov +-/ +import Mathlib.Topology.UniformSpace.Equicontinuity +import Mathlib.Topology.UniformSpace.Compact +import Mathlib.Topology.Support + +/-! +# Compact separated uniform spaces + +## Main statement + +* **Heine-Cantor** theorem: continuous functions on compact uniform spaces with values in uniform + spaces are automatically uniformly continuous. There are several variations, the main one is + `CompactSpace.uniformContinuous_of_continuous`. + +## Tags + +uniform space, uniform continuity, compact space +-/ + +open Uniformity Topology Filter UniformSpace Set + +variable {α β γ : Type*} [UniformSpace α] [UniformSpace β] + +/-! +### Heine-Cantor theorem +-/ + +/-- Heine-Cantor: a continuous function on a compact uniform space is uniformly +continuous. -/ +theorem CompactSpace.uniformContinuous_of_continuous [CompactSpace α] {f : α → β} + (h : Continuous f) : UniformContinuous f := +calc map (Prod.map f f) (𝓤 α) + = map (Prod.map f f) (𝓝ˢ (diagonal α)) := by rw [nhdsSet_diagonal_eq_uniformity] + _ ≤ 𝓝ˢ (diagonal β) := (h.prodMap h).tendsto_nhdsSet mapsTo_prod_map_diagonal + _ ≤ 𝓤 β := nhdsSet_diagonal_le_uniformity + +/-- Heine-Cantor: a continuous function on a compact set of a uniform space is uniformly +continuous. -/ +theorem IsCompact.uniformContinuousOn_of_continuous {s : Set α} {f : α → β} (hs : IsCompact s) + (hf : ContinuousOn f s) : UniformContinuousOn f s := by + rw [uniformContinuousOn_iff_restrict] + rw [isCompact_iff_compactSpace] at hs + rw [continuousOn_iff_continuous_restrict] at hf + exact CompactSpace.uniformContinuous_of_continuous hf + +/-- If `s` is compact and `f` is continuous at all points of `s`, then `f` is +"uniformly continuous at the set `s`", i.e. `f x` is close to `f y` whenever `x ∈ s` and `y` is +close to `x` (even if `y` is not itself in `s`, so this is a stronger assertion than +`UniformContinuousOn s`). -/ +theorem IsCompact.uniformContinuousAt_of_continuousAt {r : Set (β × β)} {s : Set α} + (hs : IsCompact s) (f : α → β) (hf : ∀ a ∈ s, ContinuousAt f a) (hr : r ∈ 𝓤 β) : + { x : α × α | x.1 ∈ s → (f x.1, f x.2) ∈ r } ∈ 𝓤 α := by + obtain ⟨t, ht, htsymm, htr⟩ := comp_symm_mem_uniformity_sets hr + choose U hU T hT hb using fun a ha => + exists_mem_nhds_ball_subset_of_mem_nhds ((hf a ha).preimage_mem_nhds <| mem_nhds_left _ ht) + obtain ⟨fs, hsU⟩ := hs.elim_nhds_subcover' U hU + apply mem_of_superset ((biInter_finset_mem fs).2 fun a _ => hT a a.2) + rintro ⟨a₁, a₂⟩ h h₁ + obtain ⟨a, ha, haU⟩ := Set.mem_iUnion₂.1 (hsU h₁) + apply htr + refine ⟨f a, htsymm.mk_mem_comm.1 (hb _ _ _ haU ?_), hb _ _ _ haU ?_⟩ + exacts [mem_ball_self _ (hT a a.2), mem_iInter₂.1 h a ha] + +theorem Continuous.uniformContinuous_of_tendsto_cocompact {f : α → β} {x : β} + (h_cont : Continuous f) (hx : Tendsto f (cocompact α) (𝓝 x)) : UniformContinuous f := + uniformContinuous_def.2 fun r hr => by + obtain ⟨t, ht, htsymm, htr⟩ := comp_symm_mem_uniformity_sets hr + obtain ⟨s, hs, hst⟩ := mem_cocompact.1 (hx <| mem_nhds_left _ ht) + apply + mem_of_superset + (symmetrize_mem_uniformity <| + (hs.uniformContinuousAt_of_continuousAt f fun _ _ => h_cont.continuousAt) <| + symmetrize_mem_uniformity hr) + rintro ⟨b₁, b₂⟩ h + by_cases h₁ : b₁ ∈ s; · exact (h.1 h₁).1 + by_cases h₂ : b₂ ∈ s; · exact (h.2 h₂).2 + apply htr + exact ⟨x, htsymm.mk_mem_comm.1 (hst h₁), hst h₂⟩ + +@[to_additive] +theorem HasCompactMulSupport.uniformContinuous_of_continuous {f : α → β} [One β] + (h1 : HasCompactMulSupport f) (h2 : Continuous f) : UniformContinuous f := + h2.uniformContinuous_of_tendsto_cocompact h1.is_one_at_infty + +/-- A family of functions `α → β → γ` tends uniformly to its value at `x` if `α` is locally compact, +`β` is compact and `f` is continuous on `U × (univ : Set β)` for some neighborhood `U` of `x`. -/ +theorem ContinuousOn.tendstoUniformly [LocallyCompactSpace α] [CompactSpace β] [UniformSpace γ] + {f : α → β → γ} {x : α} {U : Set α} (hxU : U ∈ 𝓝 x) (h : ContinuousOn (↿f) (U ×ˢ univ)) : + TendstoUniformly f (f x) (𝓝 x) := by + rcases LocallyCompactSpace.local_compact_nhds _ _ hxU with ⟨K, hxK, hKU, hK⟩ + have : UniformContinuousOn (↿f) (K ×ˢ univ) := + IsCompact.uniformContinuousOn_of_continuous (hK.prod isCompact_univ) + (h.mono <| prod_mono hKU Subset.rfl) + exact this.tendstoUniformly hxK + +/-- A continuous family of functions `α → β → γ` tends uniformly to its value at `x` +if `α` is weakly locally compact and `β` is compact. -/ +theorem Continuous.tendstoUniformly [WeaklyLocallyCompactSpace α] [CompactSpace β] [UniformSpace γ] + (f : α → β → γ) (h : Continuous ↿f) (x : α) : TendstoUniformly f (f x) (𝓝 x) := + let ⟨K, hK, hxK⟩ := exists_compact_mem_nhds x + have : UniformContinuousOn (↿f) (K ×ˢ univ) := + IsCompact.uniformContinuousOn_of_continuous (hK.prod isCompact_univ) h.continuousOn + this.tendstoUniformly hxK + +/-- In a product space `α × β`, assume that a function `f` is continuous on `s × k` where `k` is +compact. Then, along the fiber above any `q ∈ s`, `f` is transversely uniformly continuous, i.e., +if `p ∈ s` is close enough to `q`, then `f p x` is uniformly close to `f q x` for all `x ∈ k`. -/ +lemma IsCompact.mem_uniformity_of_prod + {α β E : Type*} [TopologicalSpace α] [TopologicalSpace β] [UniformSpace E] + {f : α → β → E} {s : Set α} {k : Set β} {q : α} {u : Set (E × E)} + (hk : IsCompact k) (hf : ContinuousOn f.uncurry (s ×ˢ k)) (hq : q ∈ s) (hu : u ∈ 𝓤 E) : + ∃ v ∈ 𝓝[s] q, ∀ p ∈ v, ∀ x ∈ k, (f p x, f q x) ∈ u := by + apply hk.induction_on (p := fun t ↦ ∃ v ∈ 𝓝[s] q, ∀ p ∈ v, ∀ x ∈ t, (f p x, f q x) ∈ u) + · exact ⟨univ, univ_mem, by simp⟩ + · intro t' t ht't ⟨v, v_mem, hv⟩ + exact ⟨v, v_mem, fun p hp x hx ↦ hv p hp x (ht't hx)⟩ + · intro t t' ⟨v, v_mem, hv⟩ ⟨v', v'_mem, hv'⟩ + refine ⟨v ∩ v', inter_mem v_mem v'_mem, fun p hp x hx ↦ ?_⟩ + rcases hx with h'x|h'x + · exact hv p hp.1 x h'x + · exact hv' p hp.2 x h'x + · rcases comp_symm_of_uniformity hu with ⟨u', u'_mem, u'_symm, hu'⟩ + intro x hx + obtain ⟨v, hv, w, hw, hvw⟩ : + ∃ v ∈ 𝓝[s] q, ∃ w ∈ 𝓝[k] x, v ×ˢ w ⊆ f.uncurry ⁻¹' {z | (f q x, z) ∈ u'} := + mem_nhdsWithin_prod_iff.1 (hf (q, x) ⟨hq, hx⟩ (mem_nhds_left (f q x) u'_mem)) + refine ⟨w, hw, v, hv, fun p hp y hy ↦ ?_⟩ + have A : (f q x, f p y) ∈ u' := hvw (⟨hp, hy⟩ : (p, y) ∈ v ×ˢ w) + have B : (f q x, f q y) ∈ u' := hvw (⟨mem_of_mem_nhdsWithin hq hv, hy⟩ : (q, y) ∈ v ×ˢ w) + exact hu' (prod_mk_mem_compRel (u'_symm A) B) + +section UniformConvergence + +/-- An equicontinuous family of functions defined on a compact uniform space is automatically +uniformly equicontinuous. -/ +theorem CompactSpace.uniformEquicontinuous_of_equicontinuous {ι : Type*} {F : ι → β → α} + [CompactSpace β] (h : Equicontinuous F) : UniformEquicontinuous F := by + rw [equicontinuous_iff_continuous] at h + rw [uniformEquicontinuous_iff_uniformContinuous] + exact CompactSpace.uniformContinuous_of_continuous h + +end UniformConvergence diff --git a/Mathlib/Topology/UniformSpace/OfCompactT2.lean b/Mathlib/Topology/UniformSpace/OfCompactT2.lean new file mode 100644 index 0000000000000..5eaf71a852f4f --- /dev/null +++ b/Mathlib/Topology/UniformSpace/OfCompactT2.lean @@ -0,0 +1,107 @@ +/- +Copyright (c) 2020 Patrick Massot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Patrick Massot, Yury Kudryashov +-/ +import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.UniformSpace.Basic + +/-! +# Compact separated uniform spaces + +## Main statement + +* `uniformSpace_of_compact_t2`: every compact T2 topological structure is induced by a uniform + structure. This uniform structure is described by `compactSpace_uniformity`. + +## Implementation notes + +The construction `uniformSpace_of_compact_t2` is not declared as an instance, as it would badly +loop. + +## Tags + +uniform space, uniform continuity, compact space +-/ + +open Uniformity Topology Filter UniformSpace Set + +variable {α β γ : Type*} [UniformSpace α] [UniformSpace β] + +/-! +### Uniformity on compact spaces +-/ + + +/-- The unique uniform structure inducing a given compact topological structure. -/ +def uniformSpaceOfCompactT2 [TopologicalSpace γ] [CompactSpace γ] [T2Space γ] : UniformSpace γ where + uniformity := 𝓝ˢ (diagonal γ) + symm := continuous_swap.tendsto_nhdsSet fun _ => Eq.symm + comp := by + /- This is the difficult part of the proof. We need to prove that, for each neighborhood `W` + of the diagonal `Δ`, there exists a smaller neighborhood `V` such that `V ○ V ⊆ W`. + -/ + set 𝓝Δ := 𝓝ˢ (diagonal γ) + -- The filter of neighborhoods of Δ + set F := 𝓝Δ.lift' fun s : Set (γ × γ) => s ○ s + -- Compositions of neighborhoods of Δ + -- If this weren't true, then there would be V ∈ 𝓝Δ such that F ⊓ 𝓟 Vᶜ ≠ ⊥ + rw [le_iff_forall_inf_principal_compl] + intro V V_in + by_contra H + haveI : NeBot (F ⊓ 𝓟 Vᶜ) := ⟨H⟩ + -- Hence compactness would give us a cluster point (x, y) for F ⊓ 𝓟 Vᶜ + obtain ⟨⟨x, y⟩, hxy⟩ : ∃ p : γ × γ, ClusterPt p (F ⊓ 𝓟 Vᶜ) := exists_clusterPt_of_compactSpace _ + -- In particular (x, y) is a cluster point of 𝓟 Vᶜ, hence is not in the interior of V, + -- and a fortiori not in Δ, so x ≠ y + have clV : ClusterPt (x, y) (𝓟 <| Vᶜ) := hxy.of_inf_right + have : (x, y) ∉ interior V := by + have : (x, y) ∈ closure Vᶜ := by rwa [mem_closure_iff_clusterPt] + rwa [closure_compl] at this + have diag_subset : diagonal γ ⊆ interior V := subset_interior_iff_mem_nhdsSet.2 V_in + have x_ne_y : x ≠ y := mt (@diag_subset (x, y)) this + -- Since γ is compact and Hausdorff, it is T₄, hence T₃. + -- So there are closed neighborhoods V₁ and V₂ of x and y contained in + -- disjoint open neighborhoods U₁ and U₂. + obtain + ⟨U₁, _, V₁, V₁_in, U₂, _, V₂, V₂_in, V₁_cl, V₂_cl, U₁_op, U₂_op, VU₁, VU₂, hU₁₂⟩ := + disjoint_nested_nhds x_ne_y + -- We set U₃ := (V₁ ∪ V₂)ᶜ so that W := U₁ ×ˢ U₁ ∪ U₂ ×ˢ U₂ ∪ U₃ ×ˢ U₃ is an open + -- neighborhood of Δ. + let U₃ := (V₁ ∪ V₂)ᶜ + have U₃_op : IsOpen U₃ := (V₁_cl.union V₂_cl).isOpen_compl + let W := U₁ ×ˢ U₁ ∪ U₂ ×ˢ U₂ ∪ U₃ ×ˢ U₃ + have W_in : W ∈ 𝓝Δ := by + rw [mem_nhdsSet_iff_forall] + rintro ⟨z, z'⟩ (rfl : z = z') + refine IsOpen.mem_nhds ?_ ?_ + · apply_rules [IsOpen.union, IsOpen.prod] + · simp only [W, mem_union, mem_prod, and_self_iff] + exact (_root_.em _).imp_left fun h => union_subset_union VU₁ VU₂ h + -- So W ○ W ∈ F by definition of F + have : W ○ W ∈ F := @mem_lift' _ _ _ (fun s => s ○ s) _ W_in + -- Porting note: was `by simpa only using mem_lift' W_in` + -- And V₁ ×ˢ V₂ ∈ 𝓝 (x, y) + have hV₁₂ : V₁ ×ˢ V₂ ∈ 𝓝 (x, y) := prod_mem_nhds V₁_in V₂_in + -- But (x, y) is also a cluster point of F so (V₁ ×ˢ V₂) ∩ (W ○ W) ≠ ∅ + -- However the construction of W implies (V₁ ×ˢ V₂) ∩ (W ○ W) = ∅. + -- Indeed assume for contradiction there is some (u, v) in the intersection. + obtain ⟨⟨u, v⟩, ⟨u_in, v_in⟩, w, huw, hwv⟩ := clusterPt_iff.mp hxy.of_inf_left hV₁₂ this + -- So u ∈ V₁, v ∈ V₂, and there exists some w such that (u, w) ∈ W and (w ,v) ∈ W. + -- Because u is in V₁ which is disjoint from U₂ and U₃, (u, w) ∈ W forces (u, w) ∈ U₁ ×ˢ U₁. + have uw_in : (u, w) ∈ U₁ ×ˢ U₁ := + (huw.resolve_right fun h => h.1 <| Or.inl u_in).resolve_right fun h => + hU₁₂.le_bot ⟨VU₁ u_in, h.1⟩ + -- Similarly, because v ∈ V₂, (w ,v) ∈ W forces (w, v) ∈ U₂ ×ˢ U₂. + have wv_in : (w, v) ∈ U₂ ×ˢ U₂ := + (hwv.resolve_right fun h => h.2 <| Or.inr v_in).resolve_left fun h => + hU₁₂.le_bot ⟨h.2, VU₂ v_in⟩ + -- Hence w ∈ U₁ ∩ U₂ which is empty. + -- So we have a contradiction + exact hU₁₂.le_bot ⟨uw_in.2, wv_in.1⟩ + nhds_eq_comap_uniformity x := by + simp_rw [nhdsSet_diagonal, comap_iSup, nhds_prod_eq, comap_prod, Function.comp_def, comap_id'] + rw [iSup_split_single _ x, comap_const_of_mem fun V => mem_of_mem_nhds] + suffices ∀ y ≠ x, comap (fun _ : γ ↦ x) (𝓝 y) ⊓ 𝓝 y ≤ 𝓝 x by simpa + intro y hxy + simp [comap_const_of_not_mem (compl_singleton_mem_nhds hxy) (not_not_intro rfl)] diff --git a/Mathlib/Topology/UniformSpace/OfFun.lean b/Mathlib/Topology/UniformSpace/OfFun.lean index 6a1ff605f190e..1fa610f447327 100644 --- a/Mathlib/Topology/UniformSpace/OfFun.lean +++ b/Mathlib/Topology/UniformSpace/OfFun.lean @@ -3,8 +3,8 @@ Copyright (c) 2023 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import Mathlib.Topology.UniformSpace.Basic import Mathlib.Algebra.Order.Monoid.Defs +import Mathlib.Topology.UniformSpace.Basic /-! # Construct a `UniformSpace` from a `dist`-like function diff --git a/Mathlib/Topology/UniformSpace/Separation.lean b/Mathlib/Topology/UniformSpace/Separation.lean index 672378e6ebef9..331e192d249dc 100644 --- a/Mathlib/Topology/UniformSpace/Separation.lean +++ b/Mathlib/Topology/UniformSpace/Separation.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Patrick Massot, Yury Kudryashov -/ import Mathlib.Tactic.ApplyFun -import Mathlib.Topology.UniformSpace.Basic import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.UniformSpace.Basic /-! # Hausdorff properties of uniform spaces. Separation quotient. diff --git a/Mathlib/Topology/UniformSpace/UniformConvergence.lean b/Mathlib/Topology/UniformSpace/UniformConvergence.lean index aca68d49538ba..48ea1c9d19316 100644 --- a/Mathlib/Topology/UniformSpace/UniformConvergence.lean +++ b/Mathlib/Topology/UniformSpace/UniformConvergence.lean @@ -3,7 +3,6 @@ Copyright (c) 2020 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import Mathlib.Topology.UniformSpace.Basic import Mathlib.Topology.UniformSpace.Cauchy /-! From ee430423941cdcdf55a1b5e84f032525cfcef151 Mon Sep 17 00:00:00 2001 From: Peiran Wu <15968905+wupr@users.noreply.github.com> Date: Tue, 19 Nov 2024 17:06:38 +0000 Subject: [PATCH 118/829] feat: define instances of `Finite`, `Fintype`, `DecidableEq` for `MulEquiv` (#17057) --- Mathlib/Algebra/Group/Equiv/Basic.lean | 5 +++++ Mathlib/Data/Finite/Prod.lean | 8 ++++++++ Mathlib/Data/Fintype/Basic.lean | 5 +++++ Mathlib/Data/Fintype/Perm.lean | 14 ++++++++++++-- 4 files changed, 30 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Group/Equiv/Basic.lean b/Mathlib/Algebra/Group/Equiv/Basic.lean index fc891d4f3cf98..81aff3106d93e 100644 --- a/Mathlib/Algebra/Group/Equiv/Basic.lean +++ b/Mathlib/Algebra/Group/Equiv/Basic.lean @@ -111,6 +111,11 @@ infixl:25 " ≃* " => MulEquiv /-- Notation for an `AddEquiv`. -/ infixl:25 " ≃+ " => AddEquiv +@[to_additive] +lemma MulEquiv.toEquiv_injective {α β : Type*} [Mul α] [Mul β] : + Function.Injective (toEquiv : (α ≃* β) → (α ≃ β)) + | ⟨_, _⟩, ⟨_, _⟩, rfl => rfl + /-- `MulEquivClass F A B` states that `F` is a type of multiplication-preserving morphisms. You should extend this class when you extend `MulEquiv`. -/ -- TODO: make this a synonym for MulHomClass? diff --git a/Mathlib/Data/Finite/Prod.lean b/Mathlib/Data/Finite/Prod.lean index 42d9d71fa3edb..c0dbebd034aaa 100644 --- a/Mathlib/Data/Finite/Prod.lean +++ b/Mathlib/Data/Finite/Prod.lean @@ -65,6 +65,14 @@ instance Equiv.finite_right {α β : Sort*} [Finite β] : Finite (α ≃ β) := instance Equiv.finite_left {α β : Sort*} [Finite α] : Finite (α ≃ β) := Finite.of_equiv _ ⟨Equiv.symm, Equiv.symm, Equiv.symm_symm, Equiv.symm_symm⟩ +@[to_additive] +instance MulEquiv.finite_left {α β : Type*} [Mul α] [Mul β] [Finite α] : Finite (α ≃* β) := + Finite.of_injective toEquiv toEquiv_injective + +@[to_additive] +instance MulEquiv.finite_right {α β : Type*} [Mul α] [Mul β] [Finite β] : Finite (α ≃* β) := + Finite.of_injective toEquiv toEquiv_injective + open Set Function variable {γ : Type*} diff --git a/Mathlib/Data/Fintype/Basic.lean b/Mathlib/Data/Fintype/Basic.lean index 90e35e73ab619..3bda18e430ecb 100644 --- a/Mathlib/Data/Fintype/Basic.lean +++ b/Mathlib/Data/Fintype/Basic.lean @@ -400,6 +400,11 @@ instance decidableEqEquivFintype [DecidableEq β] [Fintype α] : DecidableEq (α instance decidableEqEmbeddingFintype [DecidableEq β] [Fintype α] : DecidableEq (α ↪ β) := fun a b => decidable_of_iff ((a : α → β) = b) Function.Embedding.coe_injective.eq_iff +@[to_additive] +instance decidableEqMulEquivFintype {α β : Type*} [DecidableEq β] [Fintype α] [Mul α] [Mul β] : + DecidableEq (α ≃* β) := + fun a b => decidable_of_iff ((a : α → β) = b) (Injective.eq_iff DFunLike.coe_injective) + end BundledHoms instance decidableInjectiveFintype [DecidableEq α] [DecidableEq β] [Fintype α] : diff --git a/Mathlib/Data/Fintype/Perm.lean b/Mathlib/Data/Fintype/Perm.lean index 4a354e6071e5f..0c2dc9cbbd6df 100644 --- a/Mathlib/Data/Fintype/Perm.lean +++ b/Mathlib/Data/Fintype/Perm.lean @@ -140,7 +140,7 @@ theorem card_perms_of_finset : ∀ s : Finset α, #(permsOfFinset s) = (#s)! := def fintypePerm [Fintype α] : Fintype (Perm α) := ⟨permsOfFinset (@Finset.univ α _), by simp [mem_perms_of_finset_iff]⟩ -instance equivFintype [Fintype α] [Fintype β] : Fintype (α ≃ β) := +instance Equiv.instFintype [Fintype α] [Fintype β] : Fintype (α ≃ β) := if h : Fintype.card β = Fintype.card α then Trunc.recOnSubsingleton (Fintype.truncEquivFin α) fun eα => Trunc.recOnSubsingleton (Fintype.truncEquivFin β) fun eβ => @@ -148,8 +148,18 @@ instance equivFintype [Fintype α] [Fintype β] : Fintype (α ≃ β) := (equivCongr (Equiv.refl α) (eα.trans (Eq.recOn h eβ.symm)) : α ≃ α ≃ (α ≃ β)) else ⟨∅, fun x => False.elim (h (Fintype.card_eq.2 ⟨x.symm⟩))⟩ +@[deprecated (since := "2024-11-19")] alias equivFintype := Equiv.instFintype + +@[to_additive] +instance MulEquiv.instFintype + {α β : Type*} [Mul α] [Mul β] [DecidableEq α] [DecidableEq β] [Fintype α] [Fintype β] : + Fintype (α ≃* β) where + elems := Equiv.instFintype.elems.filterMap + (fun e => if h : ∀ a b : α, e (a * b) = e a * e b then (⟨e, h⟩ : α ≃* β) else none) (by aesop) + complete me := (Finset.mem_filterMap ..).mpr ⟨me.toEquiv, Finset.mem_univ _, by {simp; rfl}⟩ + theorem Fintype.card_perm [Fintype α] : Fintype.card (Perm α) = (Fintype.card α)! := - Subsingleton.elim (@fintypePerm α _ _) (@equivFintype α α _ _ _ _) ▸ card_perms_of_finset _ + Subsingleton.elim (@fintypePerm α _ _) (@Equiv.instFintype α α _ _ _ _) ▸ card_perms_of_finset _ theorem Fintype.card_equiv [Fintype α] [Fintype β] (e : α ≃ β) : Fintype.card (α ≃ β) = (Fintype.card α)! := From 22229e91bbdd842ced13a70475d39215d9429b16 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 19 Nov 2024 17:06:41 +0000 Subject: [PATCH 119/829] feat: `norm_num` extension for `Rat.num` and `Rat.den` (#19099) Closes #18826 Co-authored-by: Eric Wieser --- Mathlib/Tactic/NormNum/GCD.lean | 44 ++++++++++++++++++++++++++++++++- MathlibTest/norm_num_ext.lean | 10 ++++++++ 2 files changed, 53 insertions(+), 1 deletion(-) diff --git a/Mathlib/Tactic/NormNum/GCD.lean b/Mathlib/Tactic/NormNum/GCD.lean index 1e3b522bdd15a..7594701886b26 100644 --- a/Mathlib/Tactic/NormNum/GCD.lean +++ b/Mathlib/Tactic/NormNum/GCD.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2021 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mario Carneiro, Kyle Miller +Authors: Mario Carneiro, Kyle Miller, Eric Wieser -/ import Mathlib.Data.Int.GCD import Mathlib.Tactic.NormNum @@ -210,6 +210,48 @@ def evalIntLCM : NormNumExt where eval {u α} e := do let ⟨ed, pf⟩ := proveIntLCM ex ey return .isNat _ ed q(isInt_lcm $p $q $pf) +theorem isInt_ratNum : ∀ {q : ℚ} {n : ℤ} {n' : ℕ} {d : ℕ}, + IsRat q n d → n.natAbs = n' → n'.gcd d = 1 → IsInt q.num n + | _, n, _, d, ⟨hi, rfl⟩, rfl, h => by + constructor + have : 0 < d := Nat.pos_iff_ne_zero.mpr <| by simpa using hi.ne_zero + simp_rw [Rat.mul_num, Rat.intCast_den, invOf_eq_inv, + Rat.inv_natCast_den_of_pos this, Rat.inv_natCast_num_of_pos this, + Rat.intCast_num, one_mul, mul_one, h, Nat.cast_one, Int.ediv_one, Int.cast_id] + +theorem isNat_ratDen : ∀ {q : ℚ} {n : ℤ} {n' : ℕ} {d : ℕ}, + IsRat q n d → n.natAbs = n' → n'.gcd d = 1 → IsNat q.den d + | _, n, _, d, ⟨hi, rfl⟩, rfl, h => by + constructor + have : 0 < d := Nat.pos_iff_ne_zero.mpr <| by simpa using hi.ne_zero + simp_rw [Rat.mul_den, Rat.intCast_den, invOf_eq_inv, + Rat.inv_natCast_den_of_pos this, Rat.inv_natCast_num_of_pos this, + Rat.intCast_num, one_mul, mul_one, Nat.cast_id, h, Nat.div_one] + +/-- Evaluates the `Rat.num` function. -/ +@[nolint unusedHavesSuffices, norm_num Rat.num _] +def evalRatNum : NormNumExt where eval {u α} e := do + let .proj _ _ (q : Q(ℚ)) ← Meta.whnfR e | failure + have : u =QL 0 := ⟨⟩; have : $α =Q ℤ := ⟨⟩; have : $e =Q Rat.num $q := ⟨⟩ + let ⟨q', n, d, eq⟩ ← deriveRat q (_inst := q(inferInstance)) + let ⟨n', hn⟩ := rawIntLitNatAbs n + -- deriveRat ensures these are coprime, so the gcd will be 1 + let ⟨gcd, pf⟩ := proveNatGCD q($n') q($d) + have : $gcd =Q nat_lit 1 := ⟨⟩ + return .isInt _ n q'.num q(isInt_ratNum $eq $hn $pf) + +/-- Evaluates the `Rat.den` function. -/ +@[nolint unusedHavesSuffices, norm_num Rat.den _] +def evalRatDen : NormNumExt where eval {u α} e := do + let .proj _ _ (q : Q(ℚ)) ← Meta.whnfR e | failure + have : u =QL 0 := ⟨⟩; have : $α =Q ℕ := ⟨⟩; have : $e =Q Rat.den $q := ⟨⟩ + let ⟨q', n, d, eq⟩ ← deriveRat q (_inst := q(inferInstance)) + let ⟨n', hn⟩ := rawIntLitNatAbs n + -- deriveRat ensures these are coprime, so the gcd will be 1 + let ⟨gcd, pf⟩ := proveNatGCD q($n') q($d) + have : $gcd =Q nat_lit 1 := ⟨⟩ + return .isNat _ d q(isNat_ratDen $eq $hn $pf) + end NormNum end Tactic diff --git a/MathlibTest/norm_num_ext.lean b/MathlibTest/norm_num_ext.lean index c864ac4fd8427..de2675f268b14 100644 --- a/MathlibTest/norm_num_ext.lean +++ b/MathlibTest/norm_num_ext.lean @@ -449,3 +449,13 @@ example : (553105253 : ℤ) ∣ 553105253 * 776531401 := by norm_num1 example : ¬ (553105253 : ℤ) ∣ 553105253 * 776531401 + 1 := by norm_num1 end mod + +section num_den + +example : (6 / 15 : ℚ).num = 2 := by norm_num1 +example : (6 / 15 : ℚ).den = 5 := by norm_num1 + +example : (-6 / 15 : ℚ).num = -2 := by norm_num1 +example : (-6 / 15 : ℚ).den = 5 := by norm_num1 + +end num_den From 7242600ff92eb1a664653dc663c7572a944e37ad Mon Sep 17 00:00:00 2001 From: PieterCuijpers Date: Tue, 19 Nov 2024 17:43:39 +0000 Subject: [PATCH 120/829] feat: quantales (#17289) First definition of Quantales - a non-commutative generalization of Locales/Frames. There are still some points of discussion, which would require changes elsewhere in Mathlib, but this definition should be workeable as a starting point (I think). It's my first PR, but I've tried to stay close to what I've seen in CompleteLattices. --- Mathlib.lean | 1 + Mathlib/Algebra/Order/Quantale.lean | 179 ++++++++++++++++++++++++ Mathlib/Tactic/ToAdditive/Frontend.lean | 1 + docs/references.bib | 25 ++++ scripts/noshake.json | 6 +- 5 files changed, 209 insertions(+), 3 deletions(-) create mode 100644 Mathlib/Algebra/Order/Quantale.lean diff --git a/Mathlib.lean b/Mathlib.lean index 866e1168d05e1..7684aafa754e8 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -710,6 +710,7 @@ import Mathlib.Algebra.Order.Nonneg.Ring import Mathlib.Algebra.Order.Pi import Mathlib.Algebra.Order.Positive.Field import Mathlib.Algebra.Order.Positive.Ring +import Mathlib.Algebra.Order.Quantale import Mathlib.Algebra.Order.Rearrangement import Mathlib.Algebra.Order.Ring.Abs import Mathlib.Algebra.Order.Ring.Basic diff --git a/Mathlib/Algebra/Order/Quantale.lean b/Mathlib/Algebra/Order/Quantale.lean new file mode 100644 index 0000000000000..96bfa1624aa35 --- /dev/null +++ b/Mathlib/Algebra/Order/Quantale.lean @@ -0,0 +1,179 @@ +/- +Copyright (c) 2024 Pieter Cuijpers. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Pieter Cuijpers +-/ +import Mathlib.Algebra.Group.Defs +import Mathlib.Algebra.Order.Monoid.Unbundled.Basic +import Mathlib.Order.CompleteLattice + +/-! +# Theory of quantales + +Quantales are the non-commutative generalization of locales/frames and as such are linked +to point-free topology and order theory. Applications are found throughout logic, +quantum mechanics, and computer science (see e.g. [Vickers1989] and [Mulvey1986]). + +The most general definition of quantale occurring in literature, is that a quantale is a semigroup +distributing over a complete sup-semilattice. In our definition below, we use the fact that +every complete sup-semilattice is in fact a complete lattice, and make constructs defined for +those immediately available. Another view could be to define a quantale as a complete idempotent +semiring, i.e. a complete semiring in which + and sup coincide. However, we will often encounter +additive quantales, i.e. quantales in which the semigroup operator is thought of as addition, in +which case the link with semirings will lead to confusion notationally. + +In this file, we follow the basic definition set out on the wikipedia page on quantales, +using a mixin typeclass to make the special cases of unital, commutative, idempotent, +integral, and involutive quantales easier to add on later. + +## Main definitions + +* `IsQuantale` and `IsAddQuantale` : Typeclass mixin for a (additive) semigroup distributing + over a complete lattice, i.e satisfying `x * (sSup s) = ⨆ y ∈ s, x * y` and + `(sSup s) * y = ⨆ x ∈ s, x * y`; + +* `leftMulResiduation`, `rightMulResiduation`, `leftAddResiduation`, `rightAddResiduation` : + Defining the left- and right- residuations of the semigroup (see notation below). + +* Finally, we provide basic distributivity laws for sSup into iSup and sup, monotonicity of + the semigroup operator, and basic equivalences for left- and right-residuation. + +## Notation + +* `x ⇨ₗ y` : `sSup {z | z * x ≤ y}`, the `leftMulResiduation` of `y` over `x`; + +* `x ⇨ᵣ y` : `sSup {z | x * z ≤ y}`, the `rightMulResiduation` of `y` over `x`; + +## References + + + + + +-/ + +open Function + +/-- An additive quantale is an additive semigroup distributing over a complete lattice. -/ +class IsAddQuantale (α : Type*) [AddSemigroup α] [CompleteLattice α] where + /-- Addition is distributive over join in a quantale -/ + protected add_sSup_distrib (x : α) (s : Set α) : x + sSup s = ⨆ y ∈ s, x + y + /-- Addition is distributive over join in a quantale -/ + protected sSup_add_distrib (s : Set α) (y : α) : sSup s + y = ⨆ x ∈ s, x + y + +/-- A quantale is a semigroup distributing over a complete lattice. -/ +@[to_additive] +class IsQuantale (α : Type*) [Semigroup α] [CompleteLattice α] where + /-- Multiplication is distributive over join in a quantale -/ + protected mul_sSup_distrib (x : α) (s : Set α) : x * sSup s = ⨆ y ∈ s, x * y + /-- Multiplication is distributive over join in a quantale -/ + protected sSup_mul_distrib (s : Set α) (y : α) : sSup s * y = ⨆ x ∈ s, x * y + +section + +variable {α : Type*} {ι : Type*} {x y z : α} {s : Set α} {f : ι → α} +variable [Semigroup α] [CompleteLattice α] [IsQuantale α] + +@[to_additive] +theorem mul_sSup_distrib : x * sSup s = ⨆ y ∈ s, x * y := IsQuantale.mul_sSup_distrib _ _ + +@[to_additive] +theorem sSup_mul_distrib : sSup s * x = ⨆ y ∈ s, y * x := IsQuantale.sSup_mul_distrib _ _ + +end + +namespace IsAddQuantale + +variable {α : Type*} {ι : Type*} {x y z : α} {s : Set α} {f : ι → α} +variable [AddSemigroup α] [CompleteLattice α] [IsAddQuantale α] + +/-- Left- and right- residuation operators on an additive quantale are similar +to the Heyting operator on complete lattices, but for a non-commutative logic. +I.e. `x ≤ y ⇨ₗ z ↔ x + y ≤ z` or alternatively `x ⇨ₗ y = sSup { z | z + x ≤ y }`. -/ +def leftAddResiduation (x y : α) := sSup {z | z + x ≤ y} + +/-- Left- and right- residuation operators on an additive quantale are similar +to the Heyting operator on complete lattices, but for a non-commutative logic. +I.e. `x ≤ y ⇨ᵣ z ↔ y + x ≤ z` or alternatively `x ⇨ₗ y = sSup { z | x + z ≤ y }`." -/ +def rightAddResiduation (x y : α) := sSup {z | x + z ≤ y} + +@[inherit_doc] +scoped infixr:60 " ⇨ₗ " => leftAddResiduation + +@[inherit_doc] +scoped infixr:60 " ⇨ᵣ " => rightAddResiduation + +end IsAddQuantale + +namespace IsQuantale + +variable {α : Type*} {ι : Type*} {x y z : α} {s : Set α} {f : ι → α} +variable [Semigroup α] [CompleteLattice α] [IsQuantale α] + +/-- Left- and right-residuation operators on an additive quantale are similar to the Heyting +operator on complete lattices, but for a non-commutative logic. +I.e. `x ≤ y ⇨ₗ z ↔ x * y ≤ z` or alternatively `x ⇨ₗ y = sSup { z | z * x ≤ y }`. +-/ +@[to_additive existing] +def leftMulResiduation (x y : α) := sSup {z | z * x ≤ y} + +/-- Left- and right- residuation operators on an additive quantale are similar to the Heyting +operator on complete lattices, but for a non-commutative logic. +I.e. `x ≤ y ⇨ᵣ z ↔ y * x ≤ z` or alternatively `x ⇨ₗ y = sSup { z | x * z ≤ y }`. +-/ +@[to_additive existing] +def rightMulResiduation (x y : α) := sSup {z | x * z ≤ y} + +@[inherit_doc, to_additive existing] +scoped infixr:60 " ⇨ₗ " => leftMulResiduation + +@[inherit_doc, to_additive existing] +scoped infixr:60 " ⇨ᵣ " => rightMulResiduation + +@[to_additive] +theorem mul_iSup_distrib : x * ⨆ i, f i = ⨆ i, x * f i := by + rw [iSup, mul_sSup_distrib, iSup_range] + +@[to_additive] +theorem iSup_mul_distrib : (⨆ i, f i) * x = ⨆ i, f i * x := by + rw [iSup, sSup_mul_distrib, iSup_range] + +@[to_additive] +theorem mul_sup_distrib : x * (y ⊔ z) = (x * y) ⊔ (x * z) := by + rw [← iSup_pair, ← sSup_pair, mul_sSup_distrib] + +@[to_additive] +theorem sup_mul_distrib : (x ⊔ y) * z = (x * z) ⊔ (y * z) := by + rw [← (@iSup_pair _ _ _ (fun _? => _? * z) _ _), ← sSup_pair, sSup_mul_distrib] + +@[to_additive] +instance : MulLeftMono α where + elim := by + intro _ _ _; simp only; intro + rw [← left_eq_sup, ← mul_sup_distrib, sup_of_le_left] + trivial + +@[to_additive] +instance : MulRightMono α where + elim := by + intro _ _ _; simp only; intro + rw [← left_eq_sup, ← sup_mul_distrib, sup_of_le_left] + trivial + +@[to_additive] +theorem leftMulResiduation_le_iff_mul_le : x ≤ y ⇨ₗ z ↔ x * y ≤ z where + mp h1 := by + apply le_trans (mul_le_mul_right' h1 _) + simp_all only [leftMulResiduation, sSup_mul_distrib, Set.mem_setOf_eq, + iSup_le_iff, implies_true] + mpr h1 := le_sSup h1 + +@[to_additive] +theorem rightMulResiduation_le_iff_mul_le : x ≤ y ⇨ᵣ z ↔ y * x ≤ z where + mp h1 := by + apply le_trans (mul_le_mul_left' h1 _) + simp_all only [rightMulResiduation, mul_sSup_distrib, Set.mem_setOf_eq, + iSup_le_iff, implies_true] + mpr h1 := le_sSup h1 + +end IsQuantale diff --git a/Mathlib/Tactic/ToAdditive/Frontend.lean b/Mathlib/Tactic/ToAdditive/Frontend.lean index c7a97764f023a..25e3af975e0bf 100644 --- a/Mathlib/Tactic/ToAdditive/Frontend.lean +++ b/Mathlib/Tactic/ToAdditive/Frontend.lean @@ -964,6 +964,7 @@ def nameDict : String → List String | "powers" => ["multiples"] | "multipliable"=> ["summable"] | "gpfree" => ["apfree"] + | "quantale" => ["add", "Quantale"] | x => [x] /-- diff --git a/docs/references.bib b/docs/references.bib index 5e88d18080065..8d65ce04a4a4e 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -2792,6 +2792,22 @@ @Article{ MR577178 url = {https://doi.org/10.1007/BF00417500} } +@Article{ Mulvey1986, + author = {Mulvey, Christopher J.}, + title = {\&}, + fjournal = {Supplemento ai Rendiconti del Circolo Matem{\`a}tico di + Palermo. Serie II}, + journal = {Suppl. Rend. Circ. Mat. Palermo (2)}, + issn = {1592-9531}, + volume = {12}, + pages = {99--104}, + year = {1986}, + language = {English}, + keywords = {46L60,46L05,81Q10}, + zbmath = {4030272}, + zbl = {0633.46065} +} + @Unpublished{ Naor-2015, author = {Assaf Naor}, title = {Metric Embeddings and Lipschitz Extensions}, @@ -3422,6 +3438,15 @@ @Book{ verdier1996 mrnumber = {1453167} } +@Book{ Vickers1989, + author = {Vickers, Steven}, + title = {Topology via Logic}, + isbn = {0-521-57651-2}, + year = {1989}, + publisher = {University of Cambridge}, + language = {English} +} + @Misc{ vistoli2004, author = {Vistoli, Angelo}, title = {Notes on {Grothendieck} topologies, fibered categories and diff --git a/scripts/noshake.json b/scripts/noshake.json index 23b0f103e1de8..a1d6020a42bd4 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -323,8 +323,7 @@ ["Mathlib.Algebra.MvPolynomial.CommRing", "Mathlib.Algebra.Polynomial.Basic"], "Mathlib.RingTheory.IntegralClosure.IsIntegral.Defs": ["Mathlib.Tactic.Algebraize"], - "Mathlib.RingTheory.Finiteness.Defs": - ["Mathlib.Tactic.Algebraize"], + "Mathlib.RingTheory.Finiteness.Defs": ["Mathlib.Tactic.Algebraize"], "Mathlib.RingTheory.Binomial": ["Mathlib.Algebra.Order.Floor"], "Mathlib.RingTheory.Adjoin.Basic": ["Mathlib.LinearAlgebra.Finsupp.SumProd"], "Mathlib.RepresentationTheory.FdRep": @@ -365,7 +364,6 @@ "Mathlib.Deprecated.NatLemmas": ["Batteries.Data.Nat.Lemmas", "Batteries.WF"], "Mathlib.Deprecated.MinMax": ["Mathlib.Order.MinMax"], "Mathlib.Deprecated.ByteArray": ["Batteries.Data.ByteSubarray"], - "Mathlib.Data.ENat.Lattice": ["Mathlib.Algebra.Group.Action.Defs"], "Mathlib.Data.Vector.Basic": ["Mathlib.Control.Applicative"], "Mathlib.Data.Set.Image": ["Batteries.Tactic.Congr", "Mathlib.Data.Set.SymmDiff"], @@ -387,6 +385,7 @@ "Mathlib.Data.Int.Defs": ["Batteries.Data.Int.Order"], "Mathlib.Data.FunLike.Basic": ["Mathlib.Logic.Function.Basic"], "Mathlib.Data.Finset.Insert": ["Mathlib.Data.Finset.Attr"], + "Mathlib.Data.ENat.Lattice": ["Mathlib.Algebra.Group.Action.Defs"], "Mathlib.Data.ByteArray": ["Batteries.Data.ByteSubarray"], "Mathlib.Data.Bool.Basic": ["Batteries.Tactic.Init"], "Mathlib.Control.Traversable.Instances": ["Mathlib.Control.Applicative"], @@ -426,6 +425,7 @@ ["Mathlib.AlgebraicTopology.AlternatingFaceMapComplex"], "Mathlib.Algebra.Star.Module": ["Mathlib.Algebra.Module.LinearMap.Star"], "Mathlib.Algebra.Ring.CentroidHom": ["Mathlib.Algebra.Algebra.Defs"], + "Mathlib.Algebra.Order.Quantale": ["Mathlib.Tactic.Variable"], "Mathlib.Algebra.Order.CauSeq.Basic": ["Mathlib.Data.Setoid.Basic"], "Mathlib.Algebra.MonoidAlgebra.Basic": ["Mathlib.LinearAlgebra.Finsupp.SumProd"], From fdc34d21aae9609eaae66749df3579f64aaaa2c0 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Tue, 19 Nov 2024 17:43:40 +0000 Subject: [PATCH 121/829] feat: `IsSelfAdjoint.smul_iff` (#19216) Also generalizes some typeclass assumptions from `DivisionSemiring` to `GroupWithZero` and renames a few declarations. --- Mathlib/Algebra/Star/Basic.lean | 18 ++++--- Mathlib/Algebra/Star/Pointwise.lean | 3 +- Mathlib/Algebra/Star/SelfAdjoint.lean | 51 ++++++++++++++----- .../Analysis/InnerProductSpace/Symmetric.lean | 2 +- Mathlib/Analysis/RCLike/Basic.lean | 2 +- Mathlib/Data/Complex/Basic.lean | 2 +- 6 files changed, 53 insertions(+), 25 deletions(-) diff --git a/Mathlib/Algebra/Star/Basic.lean b/Mathlib/Algebra/Star/Basic.lean index 7e3efcae9d6b3..04f2f59e38df2 100644 --- a/Mathlib/Algebra/Star/Basic.lean +++ b/Mathlib/Algebra/Star/Basic.lean @@ -3,7 +3,6 @@ Copyright (c) 2020 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ -import Mathlib.Algebra.Field.Defs import Mathlib.Algebra.Group.Invertible.Defs import Mathlib.Algebra.GroupWithZero.Units.Lemmas import Mathlib.Algebra.Regular.Basic @@ -38,7 +37,6 @@ assert_not_exists Rat.instField universe u v w open MulOpposite -open scoped NNRat /-- Notation typeclass (with no default notation!) for an algebraic structure with a star operation. -/ @@ -365,18 +363,22 @@ open scoped ComplexConjugate end CommSemiring @[simp] -theorem star_inv' [DivisionSemiring R] [StarRing R] (x : R) : star x⁻¹ = (star x)⁻¹ := - op_injective <| (map_inv₀ (starRingEquiv : R ≃+* Rᵐᵒᵖ) x).trans (op_inv (star x)).symm +theorem star_inv₀ [GroupWithZero R] [StarMul R] (x : R) : star x⁻¹ = (star x)⁻¹ := + op_injective <| (map_inv₀ (starMulEquiv : R ≃* Rᵐᵒᵖ) x).trans (op_inv (star x)).symm + +@[deprecated (since := "2024-11-18")] alias star_inv' := star_inv₀ @[simp] -theorem star_zpow₀ [DivisionSemiring R] [StarRing R] (x : R) (z : ℤ) : star (x ^ z) = star x ^ z := - op_injective <| (map_zpow₀ (starRingEquiv : R ≃+* Rᵐᵒᵖ) x z).trans (op_zpow (star x) z).symm +theorem star_zpow₀ [GroupWithZero R] [StarMul R] (x : R) (z : ℤ) : star (x ^ z) = star x ^ z := + op_injective <| (map_zpow₀ (starMulEquiv : R ≃* Rᵐᵒᵖ) x z).trans (op_zpow (star x) z).symm /-- When multiplication is commutative, `star` preserves division. -/ @[simp] -theorem star_div' [Semifield R] [StarRing R] (x y : R) : star (x / y) = star x / star y := by +theorem star_div₀ [CommGroupWithZero R] [StarMul R] (x y : R) : star (x / y) = star x / star y := by apply op_injective - rw [division_def, op_div, mul_comm, star_mul, star_inv', op_mul, op_inv] + rw [division_def, op_div, mul_comm, star_mul, star_inv₀, op_mul, op_inv] + +@[deprecated (since := "2024-11-18")] alias star_div' := star_div₀ /-- Any commutative semiring admits the trivial `*`-structure. diff --git a/Mathlib/Algebra/Star/Pointwise.lean b/Mathlib/Algebra/Star/Pointwise.lean index 75524536ac45c..ee506f2c6f037 100644 --- a/Mathlib/Algebra/Star/Pointwise.lean +++ b/Mathlib/Algebra/Star/Pointwise.lean @@ -6,6 +6,7 @@ Authors: Jireh Loreaux import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.Star.Basic import Mathlib.Data.Set.Finite.Basic +import Mathlib.Algebra.Field.Defs /-! # Pointwise star operation on sets @@ -118,7 +119,7 @@ protected theorem star_inv [Group α] [StarMul α] (s : Set α) : s⁻¹⋆ = s protected theorem star_inv' [DivisionSemiring α] [StarRing α] (s : Set α) : s⁻¹⋆ = s⋆⁻¹ := by ext - simp only [mem_star, mem_inv, star_inv'] + simp only [mem_star, mem_inv, star_inv₀] end Set diff --git a/Mathlib/Algebra/Star/SelfAdjoint.lean b/Mathlib/Algebra/Star/SelfAdjoint.lean index ff351d4508158..db9fa87db3a50 100644 --- a/Mathlib/Algebra/Star/SelfAdjoint.lean +++ b/Mathlib/Algebra/Star/SelfAdjoint.lean @@ -235,22 +235,38 @@ protected theorem intCast (z : ℤ) : IsSelfAdjoint (z : R) := end Ring -section DivisionSemiring +section Group -variable [DivisionSemiring R] [StarRing R] +variable [Group R] [StarMul R] @[aesop safe apply] theorem inv {x : R} (hx : IsSelfAdjoint x) : IsSelfAdjoint x⁻¹ := by - simp only [isSelfAdjoint_iff, star_inv', hx.star_eq] + simp only [isSelfAdjoint_iff, star_inv, hx.star_eq] @[aesop safe apply] theorem zpow {x : R} (hx : IsSelfAdjoint x) (n : ℤ) : IsSelfAdjoint (x ^ n) := by + simp only [isSelfAdjoint_iff, star_zpow, hx.star_eq] + +end Group + +section GroupWithZero + +variable [GroupWithZero R] [StarMul R] + +@[aesop safe apply] +theorem inv₀ {x : R} (hx : IsSelfAdjoint x) : IsSelfAdjoint x⁻¹ := by + simp only [isSelfAdjoint_iff, star_inv₀, hx.star_eq] + +@[aesop safe apply] +theorem zpow₀ {x : R} (hx : IsSelfAdjoint x) (n : ℤ) : IsSelfAdjoint (x ^ n) := by simp only [isSelfAdjoint_iff, star_zpow₀, hx.star_eq] -@[simp] -protected lemma nnratCast (q : ℚ≥0) : IsSelfAdjoint (q : R) := star_nnratCast _ +end GroupWithZero -end DivisionSemiring +@[simp] +protected lemma nnratCast [DivisionSemiring R] [StarRing R] (q : ℚ≥0) : + IsSelfAdjoint (q : R) := + star_nnratCast _ section DivisionRing @@ -267,17 +283,26 @@ section Semifield variable [Semifield R] [StarRing R] theorem div {x y : R} (hx : IsSelfAdjoint x) (hy : IsSelfAdjoint y) : IsSelfAdjoint (x / y) := by - simp only [isSelfAdjoint_iff, star_div', hx.star_eq, hy.star_eq] + simp only [isSelfAdjoint_iff, star_div₀, hx.star_eq, hy.star_eq] end Semifield section SMul -variable [Star R] [AddMonoid A] [StarAddMonoid A] [SMul R A] [StarModule R A] - @[aesop safe apply] -theorem smul {r : R} (hr : IsSelfAdjoint r) {x : A} (hx : IsSelfAdjoint x) : - IsSelfAdjoint (r • x) := by simp only [isSelfAdjoint_iff, star_smul, hr.star_eq, hx.star_eq] +theorem smul [Star R] [AddMonoid A] [StarAddMonoid A] [SMul R A] [StarModule R A] + {r : R} (hr : IsSelfAdjoint r) {x : A} (hx : IsSelfAdjoint x) : + IsSelfAdjoint (r • x) := by + simp only [isSelfAdjoint_iff, star_smul, hr.star_eq, hx.star_eq] + +theorem smul_iff [Monoid R] [StarMul R] [AddMonoid A] [StarAddMonoid A] + [MulAction R A] [StarModule R A] {r : R} (hr : IsSelfAdjoint r) (hu : IsUnit r) {x : A} : + IsSelfAdjoint (r • x) ↔ IsSelfAdjoint x := by + refine ⟨fun hrx ↦ ?_, .smul hr⟩ + lift r to Rˣ using hu + rw [← inv_smul_smul r x] + replace hr : IsSelfAdjoint r := Units.ext hr.star_eq + exact hr.inv.smul hrx end SMul @@ -384,7 +409,7 @@ section Field variable [Field R] [StarRing R] instance : Inv (selfAdjoint R) where - inv x := ⟨x.val⁻¹, x.prop.inv⟩ + inv x := ⟨x.val⁻¹, x.prop.inv₀⟩ @[simp, norm_cast] theorem val_inv (x : selfAdjoint R) : ↑x⁻¹ = (x : R)⁻¹ := @@ -398,7 +423,7 @@ theorem val_div (x y : selfAdjoint R) : ↑(x / y) = (x / y : R) := rfl instance : Pow (selfAdjoint R) ℤ where - pow x z := ⟨(x : R) ^ z, x.prop.zpow z⟩ + pow x z := ⟨(x : R) ^ z, x.prop.zpow₀ z⟩ @[simp, norm_cast] theorem val_zpow (x : selfAdjoint R) (z : ℤ) : ↑(x ^ z) = (x : R) ^ z := diff --git a/Mathlib/Analysis/InnerProductSpace/Symmetric.lean b/Mathlib/Analysis/InnerProductSpace/Symmetric.lean index 2fcd08ccaa66d..fe6587c770dc5 100644 --- a/Mathlib/Analysis/InnerProductSpace/Symmetric.lean +++ b/Mathlib/Analysis/InnerProductSpace/Symmetric.lean @@ -149,7 +149,7 @@ theorem isSymmetric_iff_inner_map_self_real (T : V →ₗ[ℂ] V) : · intro h x y rw [← inner_conj_symm x (T y)] rw [inner_map_polarization T x y] - simp only [starRingEnd_apply, star_div', star_sub, star_add, star_mul] + simp only [starRingEnd_apply, star_div₀, star_sub, star_add, star_mul] simp only [← starRingEnd_apply] rw [h (x + y), h (x - y), h (x + Complex.I • y), h (x - Complex.I • y)] simp only [Complex.conj_I] diff --git a/Mathlib/Analysis/RCLike/Basic.lean b/Mathlib/Analysis/RCLike/Basic.lean index 5390b9d9b0b33..02a0ac5840750 100644 --- a/Mathlib/Analysis/RCLike/Basic.lean +++ b/Mathlib/Analysis/RCLike/Basic.lean @@ -469,7 +469,7 @@ theorem div_im (z w : K) : im (z / w) = im z * re w / normSq w - re z * im w / n @[rclike_simps] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11119): was `simp` theorem conj_inv (x : K) : conj x⁻¹ = (conj x)⁻¹ := - star_inv' _ + star_inv₀ _ lemma conj_div (x y : K) : conj (x / y) = conj x / conj y := map_div' conj conj_inv _ _ diff --git a/Mathlib/Data/Complex/Basic.lean b/Mathlib/Data/Complex/Basic.lean index 1762273d53338..e1d6851ed9e2f 100644 --- a/Mathlib/Data/Complex/Basic.lean +++ b/Mathlib/Data/Complex/Basic.lean @@ -741,7 +741,7 @@ lemma ofReal_nnqsmul (q : ℚ≥0) (r : ℝ) : ofReal (q • r) = q • r := by lemma ofReal_qsmul (q : ℚ) (r : ℝ) : ofReal (q • r) = q • r := by simp [Rat.smul_def] theorem conj_inv (x : ℂ) : conj x⁻¹ = (conj x)⁻¹ := - star_inv' _ + star_inv₀ _ @[simp, norm_cast] theorem ofReal_div (r s : ℝ) : ((r / s : ℝ) : ℂ) = r / s := map_div₀ ofRealHom r s From 6b9f5f1e22295f304afa59e7d44b3155806eb73b Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 19 Nov 2024 17:43:41 +0000 Subject: [PATCH 122/829] fix: assert_not_exists `CStarRing` that exists but not here (#19226) [Reported on Zulip](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/CI.3A.20noisy.20.22test.20mathlib.22/near/483123791) --- Mathlib/Topology/ContinuousMap/Bounded/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean b/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean index 40d279aec7e3a..81b40ea72345a 100644 --- a/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean +++ b/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean @@ -18,7 +18,7 @@ the uniform distance. -/ -assert_not_exists CStarAlgebra.star +assert_not_exists CStarRing noncomputable section From 22fb3a454a37f3bb4cb7320df78c994e575be47d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 19 Nov 2024 17:43:43 +0000 Subject: [PATCH 123/829] chore: make `Finset.subset_union_left` simp (#19249) ... and around. The corresponding `Set` lemmas already are simp From GrowthInGroups --- Mathlib/Data/Finset/Lattice/Basic.lean | 10 ++++------ 1 file changed, 4 insertions(+), 6 deletions(-) diff --git a/Mathlib/Data/Finset/Lattice/Basic.lean b/Mathlib/Data/Finset/Lattice/Basic.lean index aeb27e8cc2029..14fc97ad84371 100644 --- a/Mathlib/Data/Finset/Lattice/Basic.lean +++ b/Mathlib/Data/Finset/Lattice/Basic.lean @@ -123,9 +123,8 @@ theorem coe_union (s₁ s₂ : Finset α) : ↑(s₁ ∪ s₂) = (s₁ ∪ s₂ theorem union_subset (hs : s ⊆ u) : t ⊆ u → s ∪ t ⊆ u := sup_le <| le_iff_subset.2 hs -theorem subset_union_left {s₁ s₂ : Finset α} : s₁ ⊆ s₁ ∪ s₂ := fun _x => mem_union_left _ - -theorem subset_union_right {s₁ s₂ : Finset α} : s₂ ⊆ s₁ ∪ s₂ := fun _x => mem_union_right _ +@[simp] lemma subset_union_left : s₁ ⊆ s₁ ∪ s₂ := fun _ ↦ mem_union_left _ +@[simp] lemma subset_union_right : s₂ ⊆ s₁ ∪ s₂ := fun _ ↦ mem_union_right _ @[gcongr] theorem union_subset_union (hsu : s ⊆ u) (htv : t ⊆ v) : s ∪ t ⊆ u ∪ v := @@ -212,9 +211,8 @@ theorem mem_of_mem_inter_right {a : α} {s₁ s₂ : Finset α} (h : a ∈ s₁ theorem mem_inter_of_mem {a : α} {s₁ s₂ : Finset α} : a ∈ s₁ → a ∈ s₂ → a ∈ s₁ ∩ s₂ := and_imp.1 mem_inter.2 -theorem inter_subset_left {s₁ s₂ : Finset α} : s₁ ∩ s₂ ⊆ s₁ := fun _a => mem_of_mem_inter_left - -theorem inter_subset_right {s₁ s₂ : Finset α} : s₁ ∩ s₂ ⊆ s₂ := fun _a => mem_of_mem_inter_right +@[simp] lemma inter_subset_left : s₁ ∩ s₂ ⊆ s₁ := fun _ ↦ mem_of_mem_inter_left +@[simp] lemma inter_subset_right : s₁ ∩ s₂ ⊆ s₂ := fun _ ↦ mem_of_mem_inter_right theorem subset_inter {s₁ s₂ u : Finset α} : s₁ ⊆ s₂ → s₁ ⊆ u → s₁ ⊆ s₂ ∩ u := by simp +contextual [subset_iff, mem_inter] From 29f545b8110e796db8de688fa9e3950ce95cf678 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 19 Nov 2024 19:36:06 +0000 Subject: [PATCH 124/829] chore: generalise `div_le_div` to groups with zero (#18917) ... and lemmas around --- Archive/Imo/Imo1998Q2.lean | 2 +- Archive/Imo/Imo2001Q2.lean | 2 +- Mathlib/Algebra/GeomSum.lean | 4 +- Mathlib/Algebra/Order/Chebyshev.lean | 2 +- Mathlib/Algebra/Order/Field/Basic.lean | 69 +++++++------------ .../Order/GroupWithZero/Unbundled.lean | 64 +++++++++++++++-- Mathlib/Algebra/Order/Monovary.lean | 9 ++- .../Algebra/Polynomial/DenomsClearable.lean | 3 +- Mathlib/Analysis/Analytic/Basic.lean | 2 +- Mathlib/Analysis/Analytic/OfScalars.lean | 2 +- .../Analysis/Calculus/BumpFunction/Basic.lean | 2 +- .../Calculus/BumpFunction/Normed.lean | 2 +- Mathlib/Analysis/Calculus/MeanValue.lean | 2 +- Mathlib/Analysis/Complex/AbsMax.lean | 2 +- Mathlib/Analysis/Complex/Liouville.lean | 2 +- .../Analysis/Complex/LocallyUniformLimit.lean | 4 +- Mathlib/Analysis/Complex/Periodic.lean | 2 +- Mathlib/Analysis/Complex/Schwarz.lean | 2 +- .../Complex/UpperHalfPlane/Metric.lean | 2 +- Mathlib/Analysis/Convex/Mul.lean | 2 +- Mathlib/Analysis/Convex/Slope.lean | 16 ++--- .../Convex/SpecificFunctions/Basic.lean | 8 +-- Mathlib/Analysis/Convex/Uniform.lean | 2 +- Mathlib/Analysis/Normed/Algebra/Spectrum.lean | 2 +- .../Normed/Ring/SeminormFromBounded.lean | 10 +-- .../Normed/Ring/SeminormFromConst.lean | 6 +- Mathlib/Analysis/NormedSpace/RCLike.lean | 2 +- Mathlib/Analysis/PSeries.lean | 2 +- .../Analysis/SpecialFunctions/Log/Base.lean | 6 +- .../Analysis/SpecialFunctions/Stirling.lean | 4 +- .../SpecialFunctions/Trigonometric/Basic.lean | 6 +- Mathlib/Analysis/SpecificLimits/FloorPow.lean | 4 +- .../Additive/PluenneckeRuzsa.lean | 3 +- .../SetFamily/AhlswedeZhang.lean | 1 + Mathlib/Combinatorics/SetFamily/LYM.lean | 2 +- Mathlib/Data/Finset/Density.lean | 2 +- Mathlib/Data/NNReal/Defs.lean | 4 +- Mathlib/Data/Rat/Cast/Order.lean | 3 +- Mathlib/Data/Real/Pi/Bounds.lean | 4 +- .../RotationNumber/TranslationNumber.lean | 4 +- .../Geometry/Euclidean/Inversion/Basic.lean | 2 +- Mathlib/GroupTheory/OrderOfElement.lean | 1 + .../MeasureTheory/Integral/PeakFunction.lean | 10 +-- Mathlib/MeasureTheory/Measure/Haar/Basic.lean | 4 +- .../DiophantineApproximation.lean | 2 +- Mathlib/NumberTheory/FLT/Basic.lean | 1 + Mathlib/NumberTheory/Harmonic/Defs.lean | 1 + Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean | 2 +- .../NumberTheory/LSeries/HurwitzZetaEven.lean | 4 +- Mathlib/NumberTheory/Modular.lean | 2 +- Mathlib/NumberTheory/Padics/Hensel.lean | 2 +- Mathlib/NumberTheory/Pell.lean | 3 +- .../Transcendental/Liouville/Basic.lean | 2 +- .../Liouville/LiouvilleNumber.lean | 2 +- .../Transcendental/Liouville/Measure.lean | 4 +- Mathlib/Order/Interval/Set/IsoIoo.lean | 2 +- Mathlib/RingTheory/Multiplicity.lean | 3 +- .../RingTheory/Polynomial/Hermite/Basic.lean | 1 - Mathlib/RingTheory/PowerSeries/WellKnown.lean | 1 - Mathlib/SetTheory/Ordinal/Notation.lean | 1 + Mathlib/SetTheory/Surreal/Dyadic.lean | 2 +- Mathlib/Tactic/LinearCombination.lean | 1 + Mathlib/Tactic/LinearCombination/Lemmas.lean | 2 +- Mathlib/Tactic/NormNum/GCD.lean | 1 + Mathlib/Topology/Connected/PathConnected.lean | 2 +- Mathlib/Topology/MetricSpace/Contracting.lean | 2 +- Mathlib/Topology/TietzeExtension.lean | 2 +- MathlibTest/apply_rules.lean | 2 +- scripts/no_lints_prime_decls.txt | 2 +- scripts/noshake.json | 3 + 70 files changed, 194 insertions(+), 147 deletions(-) diff --git a/Archive/Imo/Imo1998Q2.lean b/Archive/Imo/Imo1998Q2.lean index 5e501114206ba..94378d663c520 100644 --- a/Archive/Imo/Imo1998Q2.lean +++ b/Archive/Imo/Imo1998Q2.lean @@ -202,7 +202,7 @@ end theorem clear_denominators {a b k : ℕ} (ha : 0 < a) (hb : 0 < b) : (b - 1 : ℚ) / (2 * b) ≤ k / a ↔ ((b : ℕ) - 1) * a ≤ k * (2 * b) := by - rw [div_le_div_iff] + rw [div_le_div_iff₀] -- Porting note: proof used to finish with `<;> norm_cast <;> simp [ha, hb]` · convert Nat.cast_le (α := ℚ) · aesop diff --git a/Archive/Imo/Imo2001Q2.lean b/Archive/Imo/Imo2001Q2.lean index a7e630353b751..98dc9266722ff 100644 --- a/Archive/Imo/Imo2001Q2.lean +++ b/Archive/Imo/Imo2001Q2.lean @@ -34,7 +34,7 @@ namespace Imo2001Q2 theorem bound (ha : 0 < a) (hb : 0 < b) (hc : 0 < c) : a ^ 4 / (a ^ 4 + b ^ 4 + c ^ 4) ≤ a ^ 3 / sqrt ((a ^ 3) ^ 2 + ↑8 * b ^ 3 * c ^ 3) := by - rw [div_le_div_iff (by positivity) (by positivity)] + rw [div_le_div_iff₀ (by positivity) (by positivity)] calc a ^ 4 * sqrt ((a ^ 3) ^ 2 + (8:ℝ) * b ^ 3 * c ^ 3) = a ^ 3 * (a * sqrt ((a ^ 3) ^ 2 + (8:ℝ) * b ^ 3 * c ^ 3)) := by ring _ ≤ a ^ 3 * (a ^ 4 + b ^ 4 + c ^ 4) := ?_ diff --git a/Mathlib/Algebra/GeomSum.lean b/Mathlib/Algebra/GeomSum.lean index 53093778e0883..7e6e9ce83070c 100644 --- a/Mathlib/Algebra/GeomSum.lean +++ b/Mathlib/Algebra/GeomSum.lean @@ -7,8 +7,6 @@ import Mathlib.Algebra.BigOperators.Intervals import Mathlib.Algebra.BigOperators.Ring import Mathlib.Algebra.Group.NatPowAssoc import Mathlib.Algebra.Order.BigOperators.Ring.Finset -import Mathlib.Algebra.Order.Field.Basic -import Mathlib.Algebra.Order.Ring.Abs import Mathlib.Algebra.Ring.Opposite import Mathlib.Tactic.Abel import Mathlib.Algebra.Ring.Regular @@ -382,7 +380,7 @@ theorem geom_sum_Ico_le_of_lt_one [LinearOrderedField α] {x : α} (hx : 0 ≤ x {m n : ℕ} : ∑ i ∈ Ico m n, x ^ i ≤ x ^ m / (1 - x) := by rcases le_or_lt m n with (hmn | hmn) · rw [geom_sum_Ico' h'x.ne hmn] - apply div_le_div (pow_nonneg hx _) _ (sub_pos.2 h'x) le_rfl + apply div_le_div₀ (pow_nonneg hx _) _ (sub_pos.2 h'x) le_rfl simpa using pow_nonneg hx _ · rw [Ico_eq_empty, sum_empty] · apply div_nonneg (pow_nonneg hx _) diff --git a/Mathlib/Algebra/Order/Chebyshev.lean b/Mathlib/Algebra/Order/Chebyshev.lean index 25c25f0275d3c..b0be97b492f03 100644 --- a/Mathlib/Algebra/Order/Chebyshev.lean +++ b/Mathlib/Algebra/Order/Chebyshev.lean @@ -165,6 +165,6 @@ theorem sum_div_card_sq_le_sum_sq_div_card : ((∑ i ∈ s, f i) / #s) ^ 2 ≤ (∑ i ∈ s, f i ^ 2) / #s := by obtain rfl | hs := s.eq_empty_or_nonempty · simp - rw [div_pow, div_le_div_iff (by positivity) (by positivity), sq (#s : α), mul_left_comm, + rw [div_pow, div_le_div_iff₀ (by positivity) (by positivity), sq (#s : α), mul_left_comm, ← mul_assoc] exact mul_le_mul_of_nonneg_right sq_sum_le_card_mul_sum_sq (by positivity) diff --git a/Mathlib/Algebra/Order/Field/Basic.lean b/Mathlib/Algebra/Order/Field/Basic.lean index 6d24b4b52ebbe..5832b57160d05 100644 --- a/Mathlib/Algebra/Order/Field/Basic.lean +++ b/Mathlib/Algebra/Order/Field/Basic.lean @@ -166,27 +166,6 @@ theorem one_le_inv_iff : 1 ≤ a⁻¹ ↔ 0 < a ∧ a ≤ 1 := one_le_inv_iff₀ ### Relating two divisions. -/ - -@[mono, gcongr, bound] -lemma div_le_div_of_nonneg_right (hab : a ≤ b) (hc : 0 ≤ c) : a / c ≤ b / c := by - rw [div_eq_mul_one_div a c, div_eq_mul_one_div b c] - exact mul_le_mul_of_nonneg_right hab (one_div_nonneg.2 hc) - -@[gcongr, bound] -lemma div_lt_div_of_pos_right (h : a < b) (hc : 0 < c) : a / c < b / c := by - rw [div_eq_mul_one_div a c, div_eq_mul_one_div b c] - exact mul_lt_mul_of_pos_right h (one_div_pos.2 hc) - --- Not a `mono` lemma b/c `div_le_div` is strictly more general -@[gcongr] -lemma div_le_div_of_nonneg_left (ha : 0 ≤ a) (hc : 0 < c) (h : c ≤ b) : a / b ≤ a / c := by - rw [div_eq_mul_inv, div_eq_mul_inv] - exact mul_le_mul_of_nonneg_left ((inv_le_inv₀ (hc.trans_le h) hc).mpr h) ha - -@[gcongr, bound] -lemma div_lt_div_of_pos_left (ha : 0 < a) (hc : 0 < c) (h : c < b) : a / b < a / c := by - simpa only [div_eq_mul_inv, mul_lt_mul_left ha, inv_lt_inv₀ (hc.trans h) hc] - @[deprecated (since := "2024-02-16")] alias div_le_div_of_le_of_nonneg := div_le_div_of_nonneg_right @[deprecated (since := "2024-02-16")] alias div_lt_div_of_lt := div_lt_div_of_pos_right @[deprecated (since := "2024-02-16")] alias div_le_div_of_le_left := div_le_div_of_nonneg_left @@ -196,36 +175,39 @@ lemma div_lt_div_of_pos_left (ha : 0 < a) (hc : 0 < c) (h : c < b) : a / b < a / lemma div_le_div_of_le (hc : 0 ≤ c) (hab : a ≤ b) : a / c ≤ b / c := div_le_div_of_nonneg_right hab hc -theorem div_le_div_right (hc : 0 < c) : a / c ≤ b / c ↔ a ≤ b := - ⟨le_imp_le_of_lt_imp_lt fun hab ↦ div_lt_div_of_pos_right hab hc, - fun hab ↦ div_le_div_of_nonneg_right hab hc.le⟩ +@[deprecated div_le_div_iff_of_pos_right (since := "2024-11-12")] +theorem div_le_div_right (hc : 0 < c) : a / c ≤ b / c ↔ a ≤ b := div_le_div_iff_of_pos_right hc -theorem div_lt_div_right (hc : 0 < c) : a / c < b / c ↔ a < b := - lt_iff_lt_of_le_iff_le <| div_le_div_right hc +@[deprecated div_lt_div_iff_of_pos_right (since := "2024-11-12")] +theorem div_lt_div_right (hc : 0 < c) : a / c < b / c ↔ a < b := div_lt_div_iff_of_pos_right hc -theorem div_lt_div_left (ha : 0 < a) (hb : 0 < b) (hc : 0 < c) : a / b < a / c ↔ c < b := by - simp only [div_eq_mul_inv, mul_lt_mul_left ha, inv_lt_inv₀ hb hc] +@[deprecated div_lt_div_iff_of_pos_left (since := "2024-11-13")] +theorem div_lt_div_left (ha : 0 < a) (hb : 0 < b) (hc : 0 < c) : a / b < a / c ↔ c < b := + div_lt_div_iff_of_pos_left ha hb hc +@[deprecated div_le_div_iff_of_pos_left (since := "2024-11-12")] theorem div_le_div_left (ha : 0 < a) (hb : 0 < b) (hc : 0 < c) : a / b ≤ a / c ↔ c ≤ b := - le_iff_le_iff_lt_iff_lt.2 (div_lt_div_left ha hc hb) + div_le_div_iff_of_pos_left ha hb hc -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] +@[deprecated div_lt_div_iff₀ (since := "2024-11-12")] +theorem div_lt_div_iff (b0 : 0 < b) (d0 : 0 < d) : a / b < c / d ↔ a * d < c * b := + div_lt_div_iff₀ b0 d0 -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] +@[deprecated div_le_div_iff₀ (since := "2024-11-12")] +theorem div_le_div_iff (b0 : 0 < b) (d0 : 0 < d) : a / b ≤ c / d ↔ a * d ≤ c * b := + div_le_div_iff₀ b0 d0 -@[mono, gcongr, bound] -theorem div_le_div (hc : 0 ≤ c) (hac : a ≤ c) (hd : 0 < d) (hbd : d ≤ b) : a / b ≤ c / d := by - rw [div_le_div_iff (hd.trans_le hbd) hd] - exact mul_le_mul hac hbd hd.le hc +@[deprecated div_le_div₀ (since := "2024-11-12")] +theorem div_le_div (hc : 0 ≤ c) (hac : a ≤ c) (hd : 0 < d) (hbd : d ≤ b) : a / b ≤ c / d := + div_le_div₀ hc hac hd hbd -@[gcongr] +@[deprecated div_lt_div₀ (since := "2024-11-12")] theorem div_lt_div (hac : a < c) (hbd : d ≤ b) (c0 : 0 ≤ c) (d0 : 0 < d) : a / b < c / d := - (div_lt_div_iff (d0.trans_le hbd) d0).2 (mul_lt_mul hac hbd d0 c0) + div_lt_div₀ hac hbd c0 d0 +@[deprecated div_lt_div₀' (since := "2024-11-12")] theorem div_lt_div' (hac : a ≤ c) (hbd : d < b) (c0 : 0 < c) (d0 : 0 < d) : a / b < c / d := - (div_lt_div_iff (d0.trans hbd) d0).2 (mul_lt_mul' hac hbd d0.le c0) + div_lt_div₀' hac hbd c0 d0 /-! ### Relating one division and involving `1` @@ -288,12 +270,12 @@ theorem lt_of_one_div_lt_one_div (ha : 0 < a) (h : 1 / a < 1 / b) : b < a := /-- For the single implications with fewer assumptions, see `one_div_le_one_div_of_le` and `le_of_one_div_le_one_div` -/ theorem one_div_le_one_div (ha : 0 < a) (hb : 0 < b) : 1 / a ≤ 1 / b ↔ b ≤ a := - div_le_div_left zero_lt_one ha hb + div_le_div_iff_of_pos_left zero_lt_one ha hb /-- For the single implications with fewer assumptions, see `one_div_lt_one_div_of_lt` and `lt_of_one_div_lt_one_div` -/ theorem one_div_lt_one_div (ha : 0 < a) (hb : 0 < b) : 1 / a < 1 / b ↔ b < a := - div_lt_div_left zero_lt_one ha hb + div_lt_div_iff_of_pos_left zero_lt_one ha hb theorem one_lt_one_div (h1 : 0 < a) (h2 : a < 1) : 1 < 1 / a := by rwa [lt_one_div (@zero_lt_one α _ _ _ _ _) h1, one_div_one] @@ -705,7 +687,8 @@ theorem div_two_sub_self (a : α) : a / 2 - a = -(a / 2) := by theorem add_sub_div_two_lt (h : a < b) : a + (b - a) / 2 < b := by rwa [← div_sub_div_same, sub_eq_add_neg, add_comm (b / 2), ← add_assoc, ← sub_eq_add_neg, ← - lt_sub_iff_add_lt, sub_self_div_two, sub_self_div_two, div_lt_div_right (zero_lt_two' α)] + lt_sub_iff_add_lt, sub_self_div_two, sub_self_div_two, + div_lt_div_iff_of_pos_right (zero_lt_two' α)] /-- An inequality involving `2`. -/ theorem sub_one_div_inv_le_two (a2 : 2 ≤ a) : (1 - 1 / a)⁻¹ ≤ 2 := by diff --git a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean index f11dc3c1af754..431015c2b62e0 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean @@ -1472,6 +1472,11 @@ lemma mul_inv_le_one_of_le₀ (h : a ≤ b) (hb : 0 ≤ b) : a * b⁻¹ ≤ 1 := lemma div_le_one_of_le₀ (h : a ≤ b) (hb : 0 ≤ b) : a / b ≤ 1 := div_le_of_le_mul₀ hb zero_le_one <| by rwa [one_mul] +@[mono, gcongr, bound] +lemma div_le_div_of_nonneg_right (hab : a ≤ b) (hc : 0 ≤ c) : a / c ≤ b / c := by + rw [div_eq_mul_one_div a c, div_eq_mul_one_div b c] + exact mul_le_mul_of_nonneg_right hab (one_div_nonneg.2 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₀ @@ -1498,6 +1503,12 @@ lemma le_inv_comm₀ (ha : 0 < a) (hb : 0 < b) : a ≤ b⁻¹ ↔ b ≤ a⁻¹ : lemma le_inv_of_le_inv₀ (ha : 0 < a) (h : a ≤ b⁻¹) : b ≤ a⁻¹ := (le_inv_comm₀ ha <| inv_pos.1 <| ha.trans_le h).1 h +-- Not a `mono` lemma b/c `div_le_div₀` is strictly more general +@[gcongr] +lemma div_le_div_of_nonneg_left (ha : 0 ≤ a) (hc : 0 < c) (h : c ≤ b) : a / b ≤ a / c := by + rw [div_eq_mul_inv, div_eq_mul_inv] + exact mul_le_mul_of_nonneg_left ((inv_le_inv₀ (hc.trans_le h) hc).mpr h) ha + end MulPosMono section PosMulStrictMono @@ -1599,6 +1610,11 @@ lemma div_lt_iff₀ (hc : 0 < c) : b / c < a ↔ b < a * c := by lemma inv_lt_iff_one_lt_mul₀ (ha : 0 < a) : a⁻¹ < b ↔ 1 < b * a := by rw [← mul_inv_lt_iff₀ ha, one_mul] +@[gcongr, bound] +lemma div_lt_div_of_pos_right (h : a < b) (hc : 0 < c) : a / c < b / c := by + rw [div_eq_mul_one_div a c, div_eq_mul_one_div b c] + exact mul_lt_mul_of_pos_right h (one_div_pos.2 hc) + variable [PosMulStrictMono G₀] /-- See `inv_strictAnti₀` for the implication from right-to-left with one fewer assumption. -/ @@ -1623,11 +1639,16 @@ lemma lt_inv_comm₀ (ha : 0 < a) (hb : 0 < b) : a < b⁻¹ ↔ b < a⁻¹ := by lemma lt_inv_of_lt_inv₀ (ha : 0 < a) (h : a < b⁻¹) : b < a⁻¹ := (lt_inv_comm₀ ha <| inv_pos.1 <| ha.trans h).1 h +@[gcongr, bound] +lemma div_lt_div_of_pos_left (ha : 0 < a) (hc : 0 < c) (h : c < b) : a / b < a / c := by + rw [div_eq_mul_inv, div_eq_mul_inv] + exact mul_lt_mul_of_pos_left ((inv_lt_inv₀ (hc.trans h) hc).2 h) ha + end MulPosStrictMono end PartialOrder section LinearOrder -variable [LinearOrder G₀] [ZeroLEOneClass G₀] {a b : G₀} +variable [LinearOrder G₀] [ZeroLEOneClass G₀] {a b c d : G₀} section PosMulMono variable [PosMulMono G₀] @@ -1677,6 +1698,38 @@ lemma zpow_eq_one_iff_right₀ (ha₀ : 0 ≤ a) (ha₁ : a ≠ 1) {n : ℤ} : a · exact zero_zpow_eq_one₀ simpa using zpow_right_inj₀ ha₀ ha₁ (n := 0) +variable [MulPosStrictMono G₀] + +lemma div_le_div_iff_of_pos_right (hc : 0 < c) : a / c ≤ b / c ↔ a ≤ b where + mp := le_imp_le_of_lt_imp_lt fun hab ↦ div_lt_div_of_pos_right hab hc + mpr hab := div_le_div_of_nonneg_right hab hc.le + +lemma div_lt_div_iff_of_pos_right (hc : 0 < c) : a / c < b / c ↔ a < b := + lt_iff_lt_of_le_iff_le <| div_le_div_iff_of_pos_right hc + +lemma div_lt_div_iff_of_pos_left (ha : 0 < a) (hb : 0 < b) (hc : 0 < c) : + a / b < a / c ↔ c < b := by simp only [div_eq_mul_inv, mul_lt_mul_left ha, inv_lt_inv₀ hb hc] + +lemma div_le_div_iff_of_pos_left (ha : 0 < a) (hb : 0 < b) (hc : 0 < c) : a / b ≤ a / c ↔ c ≤ b := + le_iff_le_iff_lt_iff_lt.2 (div_lt_div_iff_of_pos_left ha hc hb) + +@[mono, gcongr, bound] +lemma div_le_div₀ (hc : 0 ≤ c) (hac : a ≤ c) (hd : 0 < d) (hdb : d ≤ b) : a / b ≤ c / d := by + rw [div_eq_mul_inv, div_eq_mul_inv] + exact mul_le_mul hac ((inv_le_inv₀ (hd.trans_le hdb) hd).2 hdb) + (inv_nonneg.2 <| hd.le.trans hdb) hc + +@[gcongr] +lemma div_lt_div₀ (hac : a < c) (hdb : d ≤ b) (hc : 0 ≤ c) (hd : 0 < d) : a / b < c / d := by + rw [div_eq_mul_inv, div_eq_mul_inv] + exact mul_lt_mul hac ((inv_le_inv₀ (hd.trans_le hdb) hd).2 hdb) (inv_pos.2 <| hd.trans_le hdb) hc + +/-- See -/ +lemma div_lt_div₀' (hac : a ≤ c) (hdb : d < b) (hc : 0 < c) (hd : 0 < d) : a / b < c / d := by + rw [div_eq_mul_inv, div_eq_mul_inv] + exact mul_lt_mul' hac ((inv_lt_inv₀ (hd.trans hdb) hd).2 hdb) + (inv_nonneg.2 <| hd.le.trans hdb.le) hc + end GroupWithZero.LinearOrder section CommSemigroupHasZero @@ -1722,8 +1775,7 @@ 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 +lemma div_le_div_iff₀ (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] @@ -1751,7 +1803,7 @@ lemma div_le_comm₀ (hb : 0 < b) (hc : 0 < c) : a / b ≤ c ↔ a / c ≤ b := end PosMulMono section PosMulStrictMono -variable [PosMulStrictMono G₀] {a b c : G₀} +variable [PosMulStrictMono G₀] {a b c d : 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 @@ -1771,6 +1823,10 @@ 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] +lemma div_lt_div_iff₀ (hb : 0 < b) (hd : 0 < d) : a / b < c / d ↔ a * d < c * b := by + have := posMulStrictMono_iff_mulPosStrictMono.1 ‹_› + rw [div_lt_iff₀ hb, ← mul_div_right_comm, lt_div_iff₀ hd] + /-- 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 ‹_› diff --git a/Mathlib/Algebra/Order/Monovary.lean b/Mathlib/Algebra/Order/Monovary.lean index ebf26eae2fbc3..9ad0b00b7fa4d 100644 --- a/Mathlib/Algebra/Order/Monovary.lean +++ b/Mathlib/Algebra/Order/Monovary.lean @@ -3,7 +3,6 @@ Copyright (c) 2023 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Algebra.Order.Field.Basic import Mathlib.Algebra.Order.Group.Instances import Mathlib.Algebra.Order.Module.OrderedSMul import Mathlib.Algebra.Order.Module.Synonym @@ -291,19 +290,19 @@ alias ⟨Antivary.of_inv₀, Antivary.inv₀⟩ := antivary_inv₀ lemma MonovaryOn.div_left₀ (hf₁ : ∀ i ∈ s, 0 ≤ f₁ i) (hf₂ : ∀ i ∈ s, 0 < f₂ i) (h₁ : MonovaryOn f₁ g s) (h₂ : AntivaryOn f₂ g s) : MonovaryOn (f₁ / f₂) g s := - fun _i hi _j hj hij ↦ div_le_div (hf₁ _ hj) (h₁ hi hj hij) (hf₂ _ hj) <| h₂ hi hj hij + fun _i hi _j hj hij ↦ div_le_div₀ (hf₁ _ hj) (h₁ hi hj hij) (hf₂ _ hj) <| h₂ hi hj hij lemma AntivaryOn.div_left₀ (hf₁ : ∀ i ∈ s, 0 ≤ f₁ i) (hf₂ : ∀ i ∈ s, 0 < f₂ i) (h₁ : AntivaryOn f₁ g s) (h₂ : MonovaryOn f₂ g s) : AntivaryOn (f₁ / f₂) g s := - fun _i hi _j hj hij ↦ div_le_div (hf₁ _ hi) (h₁ hi hj hij) (hf₂ _ hi) <| h₂ hi hj hij + fun _i hi _j hj hij ↦ div_le_div₀ (hf₁ _ hi) (h₁ hi hj hij) (hf₂ _ hi) <| h₂ hi hj hij lemma Monovary.div_left₀ (hf₁ : 0 ≤ f₁) (hf₂ : StrongLT 0 f₂) (h₁ : Monovary f₁ g) (h₂ : Antivary f₂ g) : Monovary (f₁ / f₂) g := - fun _i _j hij ↦ div_le_div (hf₁ _) (h₁ hij) (hf₂ _) <| h₂ hij + fun _i _j hij ↦ div_le_div₀ (hf₁ _) (h₁ hij) (hf₂ _) <| h₂ hij lemma Antivary.div_left₀ (hf₁ : 0 ≤ f₁) (hf₂ : StrongLT 0 f₂) (h₁ : Antivary f₁ g) (h₂ : Monovary f₂ g) : Antivary (f₁ / f₂) g := - fun _i _j hij ↦ div_le_div (hf₁ _) (h₁ hij) (hf₂ _) <| h₂ hij + fun _i _j hij ↦ div_le_div₀ (hf₁ _) (h₁ hij) (hf₂ _) <| h₂ hij lemma MonovaryOn.div_right₀ (hg₁ : ∀ i ∈ s, 0 ≤ g₁ i) (hg₂ : ∀ i ∈ s, 0 < g₂ i) (h₁ : MonovaryOn f g₁ s) (h₂ : AntivaryOn f g₂ s) : MonovaryOn f (g₁ / g₂) s := diff --git a/Mathlib/Algebra/Polynomial/DenomsClearable.lean b/Mathlib/Algebra/Polynomial/DenomsClearable.lean index 3cbcf3684ebd3..1f3c6c6b23d0c 100644 --- a/Mathlib/Algebra/Polynomial/DenomsClearable.lean +++ b/Mathlib/Algebra/Polynomial/DenomsClearable.lean @@ -3,8 +3,9 @@ Copyright (c) 2020 Damiano Testa. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ -import Mathlib.Algebra.Polynomial.EraseLead import Mathlib.Algebra.Algebra.Basic +import Mathlib.Algebra.Order.Ring.Abs +import Mathlib.Algebra.Polynomial.EraseLead /-! # Denominators of evaluation of polynomials at ratios diff --git a/Mathlib/Analysis/Analytic/Basic.lean b/Mathlib/Analysis/Analytic/Basic.lean index f2a06440fe3ef..3ebd74ae9c5f3 100644 --- a/Mathlib/Analysis/Analytic/Basic.lean +++ b/Mathlib/Analysis/Analytic/Basic.lean @@ -214,7 +214,7 @@ theorem lt_radius_of_isBigO (h₀ : r ≠ 0) {a : ℝ} (ha : a ∈ Ioo (-1 : ℝ rw [← pos_iff_ne_zero, ← NNReal.coe_pos] at h₀ lift a to ℝ≥0 using ha.1.le have : (r : ℝ) < r / a := by - simpa only [div_one] using (div_lt_div_left h₀ zero_lt_one ha.1).2 ha.2 + simpa only [div_one] using (div_lt_div_iff_of_pos_left h₀ zero_lt_one ha.1).2 ha.2 norm_cast at this rw [← ENNReal.coe_lt_coe] at this refine this.trans_le (p.le_radius_of_bound C fun n => ?_) diff --git a/Mathlib/Analysis/Analytic/OfScalars.lean b/Mathlib/Analysis/Analytic/OfScalars.lean index 432b903849ad7..688d6468fb486 100644 --- a/Mathlib/Analysis/Analytic/OfScalars.lean +++ b/Mathlib/Analysis/Analytic/OfScalars.lean @@ -250,7 +250,7 @@ theorem ofScalars_radius_eq_zero_of_tendsto [NormOneClass E] cases hc' <;> aesop · filter_upwards [hc.eventually_ge_atTop (2*r⁻¹), eventually_ne_atTop 0] with n hc hn simp only [ofScalars_norm, norm_mul, norm_norm, norm_pow, NNReal.norm_eq] - rw [mul_comm ‖c n‖, ← mul_assoc, ← div_le_div_iff, mul_div_assoc] + rw [mul_comm ‖c n‖, ← mul_assoc, ← div_le_div_iff₀, mul_div_assoc] · convert hc rw [pow_succ, div_mul_cancel_left₀, NNReal.coe_inv] aesop diff --git a/Mathlib/Analysis/Calculus/BumpFunction/Basic.lean b/Mathlib/Analysis/Calculus/BumpFunction/Basic.lean index bba1643f58b06..7c20ab571383f 100644 --- a/Mathlib/Analysis/Calculus/BumpFunction/Basic.lean +++ b/Mathlib/Analysis/Calculus/BumpFunction/Basic.lean @@ -149,7 +149,7 @@ theorem support_eq : Function.support f = Metric.ball c f.rOut := by simp only [toFun, support_comp_eq_preimage, ContDiffBumpBase.support _ _ f.one_lt_rOut_div_rIn] ext x simp only [mem_ball_iff_norm, sub_zero, norm_smul, mem_preimage, Real.norm_eq_abs, abs_inv, - abs_of_pos f.rIn_pos, ← div_eq_inv_mul, div_lt_div_right f.rIn_pos] + abs_of_pos f.rIn_pos, ← div_eq_inv_mul, div_lt_div_iff_of_pos_right f.rIn_pos] theorem tsupport_eq : tsupport f = closedBall c f.rOut := by simp_rw [tsupport, f.support_eq, closure_ball _ f.rOut_pos.ne'] diff --git a/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean b/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean index fd3e53a98f231..540db4e96bee8 100644 --- a/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean +++ b/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean @@ -142,7 +142,7 @@ theorem normed_le_div_measure_closedBall_rOut [IsAddHaarMeasure μ] (K : ℝ) (h · exact f.integral_pos.le · exact f.le_one apply this.trans - rw [div_le_div_iff f.integral_pos, one_mul, ← div_le_iff₀' (pow_pos K_pos _)] + rw [div_le_div_iff₀ f.integral_pos, one_mul, ← div_le_iff₀' (pow_pos K_pos _)] · exact f.measure_closedBall_div_le_integral μ K h · exact ENNReal.toReal_pos (measure_closedBall_pos _ _ f.rOut_pos).ne' measure_closedBall_lt_top.ne diff --git a/Mathlib/Analysis/Calculus/MeanValue.lean b/Mathlib/Analysis/Calculus/MeanValue.lean index c91086f4d4c70..691da7a3e6562 100644 --- a/Mathlib/Analysis/Calculus/MeanValue.lean +++ b/Mathlib/Analysis/Calculus/MeanValue.lean @@ -117,7 +117,7 @@ theorem image_le_of_liminf_slope_right_lt_deriv_boundary' {f f' : ℝ → ℝ} { (hf'.and_eventually (HB.and (Ioc_mem_nhdsWithin_Ioi ⟨le_rfl, hy⟩))).exists refine ⟨z, ?_, hz⟩ have := (hfz.trans hzB).le - rwa [slope_def_field, slope_def_field, div_le_div_right (sub_pos.2 hz.1), hxB, + rwa [slope_def_field, slope_def_field, div_le_div_iff_of_pos_right (sub_pos.2 hz.1), hxB, sub_le_sub_iff_right] at this /-- General fencing theorem for continuous functions with an estimate on the derivative. diff --git a/Mathlib/Analysis/Complex/AbsMax.lean b/Mathlib/Analysis/Complex/AbsMax.lean index 7d73410532dd9..9f5038937c5e6 100644 --- a/Mathlib/Analysis/Complex/AbsMax.lean +++ b/Mathlib/Analysis/Complex/AbsMax.lean @@ -132,7 +132,7 @@ theorem norm_max_aux₁ [CompleteSpace F] {f : ℂ → F} {z w : ℂ} exact hz (hsub hζ) show ‖(w - z)⁻¹ • f w‖ < ‖f z‖ / r rw [norm_smul, norm_inv, norm_eq_abs, ← div_eq_inv_mul] - exact (div_lt_div_right hr).2 hw_lt + exact (div_lt_div_iff_of_pos_right hr).2 hw_lt /-! Now we drop the assumption `CompleteSpace F` by embedding `F` into its completion. diff --git a/Mathlib/Analysis/Complex/Liouville.lean b/Mathlib/Analysis/Complex/Liouville.lean index 1075a8ade0521..2d327b794e416 100644 --- a/Mathlib/Analysis/Complex/Liouville.lean +++ b/Mathlib/Analysis/Complex/Liouville.lean @@ -53,7 +53,7 @@ theorem norm_deriv_le_aux [CompleteSpace F] {c : ℂ} {R C : ℝ} {f : ℂ → F have : ∀ z ∈ sphere c R, ‖(z - c) ^ (-2 : ℤ) • f z‖ ≤ C / (R * R) := fun z (hz : abs (z - c) = R) => by simpa [-mul_inv_rev, norm_smul, hz, zpow_two, ← div_eq_inv_mul] using - (div_le_div_right (mul_pos hR hR)).2 (hC z hz) + (div_le_div_iff_of_pos_right (mul_pos hR hR)).2 (hC z hz) calc ‖deriv f c‖ = ‖(2 * π * I : ℂ)⁻¹ • ∮ z in C(c, R), (z - c) ^ (-2 : ℤ) • f z‖ := congr_arg norm (deriv_eq_smul_circleIntegral hR hf) diff --git a/Mathlib/Analysis/Complex/LocallyUniformLimit.lean b/Mathlib/Analysis/Complex/LocallyUniformLimit.lean index 14b603b86f48e..cb464f7dc6892 100644 --- a/Mathlib/Analysis/Complex/LocallyUniformLimit.lean +++ b/Mathlib/Analysis/Complex/LocallyUniformLimit.lean @@ -52,7 +52,7 @@ theorem norm_cderiv_le (hr : 0 < r) (hf : ∀ w ∈ sphere z r, ‖f w‖ ≤ M) intro w hw simp only [mem_sphere_iff_norm, norm_eq_abs] at hw simp only [norm_smul, inv_mul_eq_div, hw, norm_eq_abs, map_inv₀, Complex.abs_pow] - exact div_le_div hM (hf w hw) (sq_pos_of_pos hr) le_rfl + exact div_le_div₀ hM (hf w hw) (sq_pos_of_pos hr) le_rfl have h2 := circleIntegral.norm_integral_le_of_norm_le_const hr.le h1 simp only [cderiv, norm_smul] refine (mul_le_mul le_rfl h2 (norm_nonneg _) (norm_nonneg _)).trans (le_of_eq ?_) @@ -77,7 +77,7 @@ theorem norm_cderiv_lt (hr : 0 < r) (hfM : ∀ w ∈ sphere z r, ‖f w‖ < M) have e2 : ContinuousOn (fun w => ‖f w‖) (sphere z r) := continuous_norm.comp_continuousOn hf obtain ⟨x, hx, hx'⟩ := (isCompact_sphere z r).exists_isMaxOn e1 e2 exact ⟨‖f x‖, hfM x hx, hx'⟩ - exact (norm_cderiv_le hr hL2).trans_lt ((div_lt_div_right hr).mpr hL1) + exact (norm_cderiv_le hr hL2).trans_lt ((div_lt_div_iff_of_pos_right hr).mpr hL1) theorem norm_cderiv_sub_lt (hr : 0 < r) (hfg : ∀ w ∈ sphere z r, ‖f w - g w‖ < M) (hf : ContinuousOn f (sphere z r)) (hg : ContinuousOn g (sphere z r)) : diff --git a/Mathlib/Analysis/Complex/Periodic.lean b/Mathlib/Analysis/Complex/Periodic.lean index 2237eb83d967e..2186d2e121a79 100644 --- a/Mathlib/Analysis/Complex/Periodic.lean +++ b/Mathlib/Analysis/Complex/Periodic.lean @@ -67,7 +67,7 @@ theorem qParam_left_inv_mod_period (hh : h ≠ 0) (z : ℂ) : theorem abs_qParam_lt_iff (hh : 0 < h) (A : ℝ) (z : ℂ) : abs (qParam h z) < Real.exp (-2 * π * A / h) ↔ A < im z := by - rw [abs_qParam, Real.exp_lt_exp, div_lt_div_right hh, mul_lt_mul_left_of_neg] + rw [abs_qParam, Real.exp_lt_exp, div_lt_div_iff_of_pos_right hh, mul_lt_mul_left_of_neg] simpa using Real.pi_pos theorem qParam_tendsto (hh : 0 < h) : Tendsto (qParam h) I∞ (𝓝[≠] 0) := by diff --git a/Mathlib/Analysis/Complex/Schwarz.lean b/Mathlib/Analysis/Complex/Schwarz.lean index 9b43623115be1..91338bfc42214 100644 --- a/Mathlib/Analysis/Complex/Schwarz.lean +++ b/Mathlib/Analysis/Complex/Schwarz.lean @@ -80,7 +80,7 @@ theorem schwarz_aux {f : ℂ → ℂ} (hd : DifferentiableOn ℂ f (ball c R₁) intro z hz have hz' : z ≠ c := ne_of_mem_sphere hz hr₀.ne' rw [dslope_of_ne _ hz', slope_def_module, norm_smul, norm_inv, mem_sphere_iff_norm.1 hz, ← - div_eq_inv_mul, div_le_div_right hr₀, ← dist_eq_norm] + div_eq_inv_mul, div_le_div_iff_of_pos_right hr₀, ← dist_eq_norm] exact le_of_lt (h_maps (mem_ball.2 (by rw [mem_sphere.1 hz]; exact hr.2))) · rw [closure_ball c hr₀.ne', mem_closedBall] exact hr.1.le diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean index 79fd0dfe4e765..88a6d6642a9b9 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean @@ -79,7 +79,7 @@ protected theorem dist_comm (z w : ℍ) : dist z w = dist w z := by theorem dist_le_iff_le_sinh : dist z w ≤ r ↔ dist (z : ℂ) w / (2 * √(z.im * w.im)) ≤ sinh (r / 2) := by - rw [← div_le_div_right (zero_lt_two' ℝ), ← sinh_le_sinh, sinh_half_dist] + rw [← div_le_div_iff_of_pos_right (zero_lt_two' ℝ), ← sinh_le_sinh, sinh_half_dist] theorem dist_eq_iff_eq_sinh : dist z w = r ↔ dist (z : ℂ) w / (2 * √(z.im * w.im)) = sinh (r / 2) := by diff --git a/Mathlib/Analysis/Convex/Mul.lean b/Mathlib/Analysis/Convex/Mul.lean index 634ffe1a33097..f8a3db0f444b4 100644 --- a/Mathlib/Analysis/Convex/Mul.lean +++ b/Mathlib/Analysis/Convex/Mul.lean @@ -176,7 +176,7 @@ lemma convexOn_zpow : ∀ n : ℤ, ConvexOn 𝕜 (Ioi 0) fun x : 𝕜 ↦ x ^ n refine (convexOn_iff_forall_pos.2 ⟨convex_Ioi _, ?_⟩).pow (fun x (hx : 0 < x) ↦ by positivity) _ rintro x (hx : 0 < x) y (hy : 0 < y) a b ha hb hab field_simp - rw [div_le_div_iff, ← sub_nonneg] + rw [div_le_div_iff₀, ← sub_nonneg] · calc 0 ≤ a * b * (x - y) ^ 2 := by positivity _ = _ := by obtain rfl := eq_sub_of_add_eq hab; ring diff --git a/Mathlib/Analysis/Convex/Slope.lean b/Mathlib/Analysis/Convex/Slope.lean index 38072902d3e91..1ecc6866a18aa 100644 --- a/Mathlib/Analysis/Convex/Slope.lean +++ b/Mathlib/Analysis/Convex/Slope.lean @@ -38,7 +38,7 @@ theorem ConvexOn.slope_mono_adjacent (hf : ConvexOn 𝕜 s f) {x y z : 𝕜} (hx rw [hy] at key replace key := mul_le_mul_of_nonneg_left key hxz.le field_simp [a, b, mul_comm (z - x) _] at key ⊢ - rw [div_le_div_right] + rw [div_le_div_iff_of_pos_right] · linarith · nlinarith @@ -71,7 +71,7 @@ theorem StrictConvexOn.slope_strict_mono_adjacent (hf : StrictConvexOn 𝕜 s f) rw [hy] at key replace key := mul_lt_mul_of_pos_left key hxz field_simp [mul_comm (z - x) _] at key ⊢ - rw [div_lt_div_right] + rw [div_lt_div_iff_of_pos_right] · linarith · nlinarith @@ -100,7 +100,7 @@ theorem convexOn_of_slope_mono_adjacent (hs : Convex 𝕜 s) rw [← one_mul z, ← hab, add_mul] exact add_lt_add_right ((mul_lt_mul_left ha).2 hxz) _ have : (f y - f x) * (z - y) ≤ (f z - f y) * (y - x) := - (div_le_div_iff (sub_pos.2 hxy) (sub_pos.2 hyz)).1 (hf hx hz hxy hyz) + (div_le_div_iff₀ (sub_pos.2 hxy) (sub_pos.2 hyz)).1 (hf hx hz hxy hyz) have hxz : 0 < z - x := sub_pos.2 (hxy.trans hyz) have ha : (z - y) / (z - x) = a := by rw [eq_comm, ← sub_eq_iff_eq_add'] at hab @@ -145,7 +145,7 @@ theorem strictConvexOn_of_slope_strict_mono_adjacent (hs : Convex 𝕜 s) rw [← one_mul z, ← hab, add_mul] exact add_lt_add_right ((mul_lt_mul_left ha).2 hxz) _ have : (f y - f x) * (z - y) < (f z - f y) * (y - x) := - (div_lt_div_iff (sub_pos.2 hxy) (sub_pos.2 hyz)).1 (hf hx hz hxy hyz) + (div_lt_div_iff₀ (sub_pos.2 hxy) (sub_pos.2 hyz)).1 (hf hx hz hxy hyz) have hxz : 0 < z - x := sub_pos.2 (hxy.trans hyz) have ha : (z - y) / (z - x) = a := by rw [eq_comm, ← sub_eq_iff_eq_add'] at hab @@ -238,14 +238,14 @@ theorem ConvexOn.secant_mono_aux2 (hf : ConvexOn 𝕜 s f) {x y z : 𝕜} (hx : (hxy : x < y) (hyz : y < z) : (f y - f x) / (y - x) ≤ (f z - f x) / (z - x) := by have hxy' : 0 < y - x := by linarith have hxz' : 0 < z - x := by linarith - rw [div_le_div_iff hxy' hxz'] + rw [div_le_div_iff₀ hxy' hxz'] linarith only [hf.secant_mono_aux1 hx hz hxy hyz] theorem ConvexOn.secant_mono_aux3 (hf : ConvexOn 𝕜 s f) {x y z : 𝕜} (hx : x ∈ s) (hz : z ∈ s) (hxy : x < y) (hyz : y < z) : (f z - f x) / (z - x) ≤ (f z - f y) / (z - y) := by have hyz' : 0 < z - y := by linarith have hxz' : 0 < z - x := by linarith - rw [div_le_div_iff hxz' hyz'] + rw [div_le_div_iff₀ hxz' hyz'] linarith only [hf.secant_mono_aux1 hx hz hxy hyz] /-- If `f : 𝕜 → 𝕜` is convex, then for any point `a` the slope of the secant line of `f` through `a` @@ -284,14 +284,14 @@ theorem StrictConvexOn.secant_strict_mono_aux2 (hf : StrictConvexOn 𝕜 s f) {x (hz : z ∈ s) (hxy : x < y) (hyz : y < z) : (f y - f x) / (y - x) < (f z - f x) / (z - x) := by have hxy' : 0 < y - x := by linarith have hxz' : 0 < z - x := by linarith - rw [div_lt_div_iff hxy' hxz'] + rw [div_lt_div_iff₀ hxy' hxz'] linarith only [hf.secant_strict_mono_aux1 hx hz hxy hyz] theorem StrictConvexOn.secant_strict_mono_aux3 (hf : StrictConvexOn 𝕜 s f) {x y z : 𝕜} (hx : x ∈ s) (hz : z ∈ s) (hxy : x < y) (hyz : y < z) : (f z - f x) / (z - x) < (f z - f y) / (z - y) := by have hyz' : 0 < z - y := by linarith have hxz' : 0 < z - x := by linarith - rw [div_lt_div_iff hxz' hyz'] + rw [div_lt_div_iff₀ hxz' hyz'] linarith only [hf.secant_strict_mono_aux1 hx hz hxy hyz] /-- If `f : 𝕜 → 𝕜` is strictly convex, then for any point `a` the slope of the secant line of `f` diff --git a/Mathlib/Analysis/Convex/SpecificFunctions/Basic.lean b/Mathlib/Analysis/Convex/SpecificFunctions/Basic.lean index b293f4c6a7eb3..e58b1bbfd2213 100644 --- a/Mathlib/Analysis/Convex/SpecificFunctions/Basic.lean +++ b/Mathlib/Analysis/Convex/SpecificFunctions/Basic.lean @@ -110,7 +110,7 @@ theorem one_add_mul_self_lt_rpow_one_add {s : ℝ} (hs : -1 ≤ s) (hs' : s ≠ · 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_iff_of_pos_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] @@ -149,7 +149,7 @@ theorem rpow_one_add_lt_one_add_mul_self {s : ℝ} (hs : -1 ≤ s) (hs' : s ≠ · 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_iff_of_pos_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,7 +175,7 @@ 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', + rw [div_lt_iff₀ q, ← div_lt_div_iff_of_pos_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, @@ -186,7 +186,7 @@ theorem strictConvexOn_rpow {p : ℝ} (hp : 1 < p) : StrictConvexOn ℝ (Ici 0) · 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', + rw [lt_div_iff₀ q, ← div_lt_div_iff_of_pos_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'] diff --git a/Mathlib/Analysis/Convex/Uniform.lean b/Mathlib/Analysis/Convex/Uniform.lean index 152127c0a4736..cbea3a0fbb080 100644 --- a/Mathlib/Analysis/Convex/Uniform.lean +++ b/Mathlib/Analysis/Convex/Uniform.lean @@ -118,7 +118,7 @@ theorem exists_forall_closed_ball_dist_add_le_two_mul_sub (hε : 0 < ε) (r : rw [← div_le_one hr, div_eq_inv_mul, ← norm_smul_of_nonneg (inv_nonneg.2 hr.le)] at hx hy have := h hx hy simp_rw [← smul_add, ← smul_sub, norm_smul_of_nonneg (inv_nonneg.2 hr.le), ← div_eq_inv_mul, - div_le_div_right hr, div_le_iff₀ hr, sub_mul] at this + div_le_div_iff_of_pos_right hr, div_le_iff₀ hr, sub_mul] at this exact this hxy end SeminormedAddCommGroup diff --git a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean index 5dd3a9d8f7ddd..4f081227992a4 100644 --- a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean +++ b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean @@ -483,7 +483,7 @@ theorem exp_mem_exp [RCLike 𝕜] [NormedRing A] [NormedAlgebra 𝕜 A] [Complet refine .of_norm_bounded_eventually _ (Real.summable_pow_div_factorial ‖a - ↑ₐ z‖) ?_ filter_upwards [Filter.eventually_cofinite_ne 0] with n hn rw [norm_smul, mul_comm, norm_inv, RCLike.norm_natCast, ← div_eq_mul_inv] - exact div_le_div (pow_nonneg (norm_nonneg _) n) (norm_pow_le' (a - ↑ₐ z) (zero_lt_iff.mpr hn)) + exact div_le_div₀ (pow_nonneg (norm_nonneg _) n) (norm_pow_le' (a - ↑ₐ z) (zero_lt_iff.mpr hn)) (mod_cast Nat.factorial_pos n) (mod_cast Nat.factorial_le (lt_add_one n).le) have h₀ : (∑' n : ℕ, ((n + 1).factorial⁻¹ : 𝕜) • (a - ↑ₐ z) ^ (n + 1)) = (a - ↑ₐ z) * b := by simpa only [mul_smul_comm, pow_succ'] using hb.tsum_mul_left (a - ↑ₐ z) diff --git a/Mathlib/Analysis/Normed/Ring/SeminormFromBounded.lean b/Mathlib/Analysis/Normed/Ring/SeminormFromBounded.lean index a98f45db8cc03..4aeac560e1e53 100644 --- a/Mathlib/Analysis/Normed/Ring/SeminormFromBounded.lean +++ b/Mathlib/Analysis/Normed/Ring/SeminormFromBounded.lean @@ -189,13 +189,13 @@ theorem seminormFromBounded_mul (f_nonneg : 0 ≤ f) simp_rw [mul_comm, hxyz, zero_div] exact div_nonneg (mul_nonneg (seminormFromBounded_nonneg f_nonneg f_mul y) (f_nonneg _)) (f_nonneg _) - · rw [div_le_div_right (lt_of_le_of_ne' (f_nonneg _) hz), mul_comm (f (x * z))] + · rw [div_le_div_iff_of_pos_right (lt_of_le_of_ne' (f_nonneg _) hz), mul_comm (f (x * z))] by_cases hxz : f (x * z) = 0 · rw [mul_comm x y, mul_assoc, mul_comm y, map_mul_zero_of_map_zero f_nonneg f_mul hxz y] exact mul_nonneg (seminormFromBounded_nonneg f_nonneg f_mul y) (f_nonneg _) · rw [← div_le_iff₀ (lt_of_le_of_ne' (f_nonneg _) hxz)] apply le_ciSup_of_le (seminormFromBounded_bddAbove_range f_nonneg f_mul y) (x * z) - rw [div_le_div_right (lt_of_le_of_ne' (f_nonneg _) hxz), mul_comm x y, mul_assoc] + rw [div_le_div_iff_of_pos_right (lt_of_le_of_ne' (f_nonneg _) hxz), mul_comm x y, mul_assoc] /-- If `f : R → ℝ` is a nonzero, nonnegative, multiplicatively bounded function, then `seminormFromBounded' f 1 = 1`. -/ @@ -242,7 +242,7 @@ theorem seminormFromBounded_add (f_nonneg : 0 ≤ f) (le_ciSup_of_le (seminormFromBounded_bddAbove_range f_nonneg f_mul y) z (le_refl _))) by_cases hz : f z = 0 · simp only [hz, div_zero, zero_add, le_refl, or_self_iff] - · rw [div_add_div_same, div_le_div_right (lt_of_le_of_ne' (f_nonneg _) hz), add_mul] + · rw [div_add_div_same, div_le_div_iff_of_pos_right (lt_of_le_of_ne' (f_nonneg _) hz), add_mul] exact f_add _ _ /-- `seminormFromBounded'` is a ring seminorm on `R`. -/ @@ -268,8 +268,8 @@ theorem seminormFromBounded_isNonarchimedean (f_nonneg : 0 ≤ f) · exact Or.inr <| le_ciSup_of_le (seminormFromBounded_bddAbove_range f_nonneg f_mul y) z hfy by_cases hz : f z = 0 · simp only [hz, div_zero, le_refl, or_self_iff] - · rw [div_le_div_right (lt_of_le_of_ne' (f_nonneg _) hz), - div_le_div_right (lt_of_le_of_ne' (f_nonneg _) hz), add_mul, ← le_max_iff] + · rw [div_le_div_iff_of_pos_right (lt_of_le_of_ne' (f_nonneg _) hz), + div_le_div_iff_of_pos_right (lt_of_le_of_ne' (f_nonneg _) hz), add_mul, ← le_max_iff] exact hna _ _ /-- If `f : R → ℝ` is a nonnegative, multiplicatively bounded function and `x : R` is diff --git a/Mathlib/Analysis/Normed/Ring/SeminormFromConst.lean b/Mathlib/Analysis/Normed/Ring/SeminormFromConst.lean index c45e2c357f115..91ff0ab51079c 100644 --- a/Mathlib/Analysis/Normed/Ring/SeminormFromConst.lean +++ b/Mathlib/Analysis/Normed/Ring/SeminormFromConst.lean @@ -89,10 +89,10 @@ theorem seminormFromConst_seq_antitone (x : R) : Antitone (seminormFromConst_seq nth_rw 1 [← Nat.add_sub_of_le hmn] rw [pow_add, ← mul_assoc] have hc_pos : 0 < f c := lt_of_le_of_ne (apply_nonneg f _) hc.symm - apply le_trans ((div_le_div_right (pow_pos hc_pos _)).mpr (map_mul_le_mul f _ _)) + apply le_trans ((div_le_div_iff_of_pos_right (pow_pos hc_pos _)).mpr (map_mul_le_mul f _ _)) by_cases heq : m = n · have hnm : n - m = 0 := by rw [heq, Nat.sub_self n] - rw [hnm, heq, div_le_div_right (pow_pos hc_pos _), pow_zero] + rw [hnm, heq, div_le_div_iff_of_pos_right (pow_pos hc_pos _), pow_zero] conv_rhs => rw [← mul_one (f (x * c ^ n))] exact mul_le_mul_of_nonneg_left hf1 (apply_nonneg f _) · have h1 : 1 ≤ n - m := by @@ -150,7 +150,7 @@ def seminormFromConst : RingSeminorm R where (seminormFromConst_isLimit hf1 hc hpm y)) (fun n ↦ ?_) simp only [seminormFromConst_seq] rw [div_mul_div_comm, ← pow_add, two_mul, - div_le_div_right (pow_pos (lt_of_le_of_ne (apply_nonneg f _) hc.symm) _), pow_add, + div_le_div_iff_of_pos_right (pow_pos (lt_of_le_of_ne (apply_nonneg f _) hc.symm) _), pow_add, ← mul_assoc, mul_comm (x * y), ← mul_assoc, mul_assoc, mul_comm (c ^ n)] exact map_mul_le_mul f (x * c ^ n) (y * c ^ n) diff --git a/Mathlib/Analysis/NormedSpace/RCLike.lean b/Mathlib/Analysis/NormedSpace/RCLike.lean index 6b6b64a156b06..f578908f625cc 100644 --- a/Mathlib/Analysis/NormedSpace/RCLike.lean +++ b/Mathlib/Analysis/NormedSpace/RCLike.lean @@ -65,7 +65,7 @@ theorem LinearMap.bound_of_sphere_bound {r : ℝ} (r_pos : 0 < r) (c : ℝ) (f : simp only [z_zero, RCLike.ofReal_eq_zero, norm_eq_zero, Ne, not_false_iff] rw [eq, norm_mul, norm_div, RCLike.norm_coe_norm, RCLike.norm_of_nonneg r_pos.le, div_mul_eq_mul_div, div_mul_eq_mul_div, mul_comm] - apply div_le_div _ _ r_pos rfl.ge + apply div_le_div₀ _ _ r_pos rfl.ge · exact mul_nonneg ((norm_nonneg _).trans norm_f_z₁) (norm_nonneg z) apply mul_le_mul norm_f_z₁ rfl.le (norm_nonneg z) ((norm_nonneg _).trans norm_f_z₁) diff --git a/Mathlib/Analysis/PSeries.lean b/Mathlib/Analysis/PSeries.lean index 0dc1b5b99f4d7..b077709ddf4da 100644 --- a/Mathlib/Analysis/PSeries.lean +++ b/Mathlib/Analysis/PSeries.lean @@ -388,7 +388,7 @@ theorem sum_Ioc_inv_sq_le_sub {k n : ℕ} (hk : k ≠ 0) (h : k ≤ n) : have A : 0 < (n : α) := by simpa using hk.bot_lt.trans_le hn have B : 0 < (n : α) + 1 := by linarith field_simp - rw [div_le_div_iff _ A, ← sub_nonneg] + rw [div_le_div_iff₀ _ A, ← sub_nonneg] · ring_nf rw [add_comm] exact B.le diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Base.lean b/Mathlib/Analysis/SpecialFunctions/Log/Base.lean index 1adccc79ae6df..5e688563f5ae7 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Base.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Base.lean @@ -177,7 +177,7 @@ private theorem b_ne_one' : b ≠ 1 := by linarith @[simp] theorem logb_le_logb (h : 0 < x) (h₁ : 0 < y) : logb b x ≤ logb b y ↔ x ≤ y := by - rw [logb, logb, div_le_div_right (log_pos hb), log_le_log_iff h h₁] + rw [logb, logb, div_le_div_iff_of_pos_right (log_pos hb), log_le_log_iff h h₁] @[gcongr] theorem logb_le_logb_of_le (h : 0 < x) (hxy : x ≤ y) : logb b x ≤ logb b y := @@ -185,12 +185,12 @@ theorem logb_le_logb_of_le (h : 0 < x) (hxy : x ≤ y) : logb b x ≤ logb b y : @[gcongr] theorem logb_lt_logb (hx : 0 < x) (hxy : x < y) : logb b x < logb b y := by - rw [logb, logb, div_lt_div_right (log_pos hb)] + rw [logb, logb, div_lt_div_iff_of_pos_right (log_pos hb)] exact log_lt_log hx hxy @[simp] theorem logb_lt_logb_iff (hx : 0 < x) (hy : 0 < y) : logb b x < logb b y ↔ x < y := by - rw [logb, logb, div_lt_div_right (log_pos hb)] + rw [logb, logb, div_lt_div_iff_of_pos_right (log_pos hb)] exact log_lt_log_iff hx hy theorem logb_le_iff_le_rpow (hx : 0 < x) : logb b x ≤ y ↔ x ≤ b ^ y := by diff --git a/Mathlib/Analysis/SpecialFunctions/Stirling.lean b/Mathlib/Analysis/SpecialFunctions/Stirling.lean index 35d9ac4c47d5b..c42129b43cc42 100644 --- a/Mathlib/Analysis/SpecialFunctions/Stirling.lean +++ b/Mathlib/Analysis/SpecialFunctions/Stirling.lean @@ -122,9 +122,9 @@ theorem log_stirlingSeq_sub_log_stirlingSeq_succ (n : ℕ) : convert H using 1 <;> field_simp [h₃.ne'] refine (log_stirlingSeq_diff_le_geo_sum n).trans ?_ push_cast - rw [div_le_div_iff h₂ h₁] + rw [div_le_div_iff₀ h₂ h₁] field_simp [h₃.ne'] - rw [div_le_div_right h₃] + rw [div_le_div_iff_of_pos_right h₃] ring_nf norm_cast omega diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean index 4c002f27c027e..68908a896942d 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean @@ -137,11 +137,11 @@ theorem pi_div_two_le_two : π / 2 ≤ 2 := by exact (Classical.choose_spec exists_cos_eq_zero).1.2 theorem two_le_pi : (2 : ℝ) ≤ π := - (div_le_div_right (show (0 : ℝ) < 2 by norm_num)).1 + (div_le_div_iff_of_pos_right (show (0 : ℝ) < 2 by norm_num)).1 (by rw [div_self (two_ne_zero' ℝ)]; exact one_le_pi_div_two) theorem pi_le_four : π ≤ 4 := - (div_le_div_right (show (0 : ℝ) < 2 by norm_num)).1 + (div_le_div_iff_of_pos_right (show (0 : ℝ) < 2 by norm_num)).1 (calc π / 2 ≤ 2 := pi_div_two_le_two _ = 4 / 2 := by norm_num) @@ -887,7 +887,7 @@ theorem tan_nonpos_of_nonpos_of_neg_pi_div_two_le {x : ℝ} (hx0 : x ≤ 0) (hpx theorem strictMonoOn_tan : StrictMonoOn tan (Ioo (-(π / 2)) (π / 2)) := by rintro x hx y hy hlt rw [tan_eq_sin_div_cos, tan_eq_sin_div_cos, - div_lt_div_iff (cos_pos_of_mem_Ioo hx) (cos_pos_of_mem_Ioo hy), mul_comm, ← sub_pos, ← sin_sub] + div_lt_div_iff₀ (cos_pos_of_mem_Ioo hx) (cos_pos_of_mem_Ioo hy), mul_comm, ← sub_pos, ← sin_sub] exact sin_pos_of_pos_of_lt_pi (sub_pos.2 hlt) <| by linarith [hx.1, hy.2] theorem tan_lt_tan_of_lt_of_lt_pi_div_two {x y : ℝ} (hx₁ : -(π / 2) < x) (hy₂ : y < π / 2) diff --git a/Mathlib/Analysis/SpecificLimits/FloorPow.lean b/Mathlib/Analysis/SpecificLimits/FloorPow.lean index 819d71a7534da..d03aabb125742 100644 --- a/Mathlib/Analysis/SpecificLimits/FloorPow.lean +++ b/Mathlib/Analysis/SpecificLimits/FloorPow.lean @@ -220,7 +220,7 @@ theorem sum_div_pow_sq_le_div_sq (N : ℕ) {j : ℝ} (hj : 0 < j) {c : ℝ} (hc have cpos : 0 < c := zero_lt_one.trans hc have A : (0 : ℝ) < c⁻¹ ^ 2 := sq_pos_of_pos (inv_pos.2 cpos) have B : c ^ 2 * ((1 : ℝ) - c⁻¹ ^ 2)⁻¹ ≤ c ^ 3 * (c - 1)⁻¹ := by - rw [← div_eq_mul_inv, ← div_eq_mul_inv, div_le_div_iff _ (sub_pos.2 hc)] + 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_of_one_lt₀ hc) two_ne_zero) have : c ^ 3 = c ^ 2 * c := by ring @@ -287,7 +287,7 @@ theorem sum_div_nat_floor_pow_sq_le_div_sq (N : ℕ) {j : ℝ} (hj : 0 < j) {c : exact fun k hk ↦ hk.trans_le <| Nat.floor_le (by positivity) _ ≤ ∑ i ∈ range N with j < c ^ i, (1 - c⁻¹)⁻¹ ^ 2 * ((1 : ℝ) / (c ^ i) ^ 2) := by refine sum_le_sum fun i _hi => ?_ - rw [mul_div_assoc', mul_one, div_le_div_iff]; rotate_left + 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₀, hc.le, Nat.one_le_cast, Nat.cast_one] diff --git a/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean b/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean index 286aa6a748287..a65c2df01ec20 100644 --- a/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean +++ b/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, George Shakan -/ import Mathlib.Algebra.Group.Pointwise.Finset.Basic -import Mathlib.Algebra.Order.Field.Basic import Mathlib.Algebra.Order.Field.Rat import Mathlib.Algebra.Order.Ring.Basic import Mathlib.Combinatorics.Enumerative.DoubleCounting @@ -164,7 +163,7 @@ private theorem mul_aux (hA : A.Nonempty) (hAB : A ⊆ B) have hA₀ : (0 : ℚ≥0) < #A := cast_pos.2 hA.card_pos have hA₀' : (0 : ℚ≥0) < #A' := cast_pos.2 hA'.card_pos exact mod_cast - (div_le_div_iff hA₀ hA₀').1 + (div_le_div_iff₀ hA₀ hA₀').1 (h _ <| mem_erase_of_ne_of_mem hA'.ne_empty <| mem_powerset.2 <| hAA'.trans hAB) /-- **Ruzsa's triangle inequality**. Multiplication version. -/ diff --git a/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean b/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean index d2173e136b147..be6ce35efc01e 100644 --- a/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean +++ b/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean @@ -6,6 +6,7 @@ Authors: Yaël Dillies, Vladimir Ivanov import Mathlib.Algebra.BigOperators.Intervals import Mathlib.Algebra.BigOperators.Ring import Mathlib.Algebra.Order.BigOperators.Group.Finset +import Mathlib.Algebra.Order.Field.Basic import Mathlib.Data.Finset.Sups import Mathlib.Tactic.FieldSimp import Mathlib.Tactic.Ring diff --git a/Mathlib/Combinatorics/SetFamily/LYM.lean b/Mathlib/Combinatorics/SetFamily/LYM.lean index 4f771b1d44162..67a603abd4165 100644 --- a/Mathlib/Combinatorics/SetFamily/LYM.lean +++ b/Mathlib/Combinatorics/SetFamily/LYM.lean @@ -93,7 +93,7 @@ theorem card_div_choose_le_card_shadow_div_choose (hr : r ≠ 0) · rw [choose_eq_zero_of_lt hr', cast_zero, div_zero] exact div_nonneg (cast_nonneg _) (cast_nonneg _) replace h𝒜 := card_mul_le_card_shadow_mul h𝒜 - rw [div_le_div_iff] <;> norm_cast + rw [div_le_div_iff₀] <;> norm_cast · cases' r with r · exact (hr rfl).elim rw [tsub_add_eq_add_tsub hr', add_tsub_add_eq_tsub_right] at h𝒜 diff --git a/Mathlib/Data/Finset/Density.lean b/Mathlib/Data/Finset/Density.lean index b8fcbb8dcf9b9..b8d5585f00ed8 100644 --- a/Mathlib/Data/Finset/Density.lean +++ b/Mathlib/Data/Finset/Density.lean @@ -3,11 +3,11 @@ Copyright (c) 2023 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Algebra.Order.Field.Basic import Mathlib.Algebra.Order.Field.Rat import Mathlib.Data.Fintype.Card import Mathlib.Data.NNRat.Order import Mathlib.Data.Rat.Cast.CharZero +import Mathlib.Tactic.Positivity /-! # Density of a finite set diff --git a/Mathlib/Data/NNReal/Defs.lean b/Mathlib/Data/NNReal/Defs.lean index ae12ffd52c40f..210cfbe55782c 100644 --- a/Mathlib/Data/NNReal/Defs.lean +++ b/Mathlib/Data/NNReal/Defs.lean @@ -791,13 +791,15 @@ theorem lt_div_iff' {a b r : ℝ≥0} (hr : r ≠ 0) : a < b / r ↔ r * a < b : theorem mul_lt_of_lt_div {a b r : ℝ≥0} (h : a < b / r) : a * r < b := (lt_div_iff₀ <| pos_iff_ne_zero.2 fun hr => False.elim <| by simp [hr] at h).1 h +@[deprecated div_le_div_of_nonneg_left (since := "2024-11-12")] theorem div_le_div_left_of_le {a b c : ℝ≥0} (c0 : c ≠ 0) (cb : c ≤ b) : a / b ≤ a / c := div_le_div_of_nonneg_left (zero_le _) c0.bot_lt cb +@[deprecated div_le_div_iff_of_pos_left (since := "2024-11-12")] nonrec theorem div_le_div_left {a b c : ℝ≥0} (a0 : 0 < a) (b0 : 0 < b) (c0 : 0 < c) : a / b ≤ a / c ↔ c ≤ b := - div_le_div_left a0 b0 c0 + div_le_div_iff_of_pos_left a0 b0 c0 theorem le_of_forall_lt_one_mul_le {x y : ℝ≥0} (h : ∀ a < 1, a * x ≤ y) : x ≤ y := le_of_forall_ge_of_dense fun a ha => by diff --git a/Mathlib/Data/Rat/Cast/Order.lean b/Mathlib/Data/Rat/Cast/Order.lean index 637675c471ea0..618e76fd32375 100644 --- a/Mathlib/Data/Rat/Cast/Order.lean +++ b/Mathlib/Data/Rat/Cast/Order.lean @@ -6,7 +6,6 @@ Authors: Johannes Hölzl, Mario Carneiro import Mathlib.Algebra.Order.Field.Rat import Mathlib.Data.Rat.Cast.CharZero import Mathlib.Tactic.Positivity.Core -import Mathlib.Algebra.Order.Field.Basic /-! # Casts of rational numbers into linear ordered fields. @@ -147,7 +146,7 @@ namespace NNRat variable {K} [LinearOrderedSemifield K] {p q : ℚ≥0} theorem cast_strictMono : StrictMono ((↑) : ℚ≥0 → K) := fun p q h => by - rwa [NNRat.cast_def, NNRat.cast_def, div_lt_div_iff, ← Nat.cast_mul, ← Nat.cast_mul, + rwa [NNRat.cast_def, NNRat.cast_def, div_lt_div_iff₀, ← Nat.cast_mul, ← Nat.cast_mul, Nat.cast_lt (α := K), ← NNRat.lt_def] · simp · simp diff --git a/Mathlib/Data/Real/Pi/Bounds.lean b/Mathlib/Data/Real/Pi/Bounds.lean index d04c703754a53..14e62e0f82c7c 100644 --- a/Mathlib/Data/Real/Pi/Bounds.lean +++ b/Mathlib/Data/Real/Pi/Bounds.lean @@ -68,7 +68,7 @@ theorem sqrtTwoAddSeries_step_up (c d : ℕ) {a b n : ℕ} {z : ℝ} (hz : sqrtT have hb' : 0 < (b : ℝ) := Nat.cast_pos.2 hb have hd' : 0 < (d : ℝ) := Nat.cast_pos.2 hd rw [sqrt_le_left (div_nonneg c.cast_nonneg d.cast_nonneg), div_pow, - add_div_eq_mul_add_div _ _ (ne_of_gt hb'), div_le_div_iff hb' (pow_pos hd' _)] + add_div_eq_mul_add_div _ _ (ne_of_gt hb'), div_le_div_iff₀ hb' (pow_pos hd' _)] exact mod_cast h /-- From a lower bound on `sqrtTwoAddSeries 0 n = 2 cos (π / 2 ^ (n+1))` of the form @@ -91,7 +91,7 @@ theorem sqrtTwoAddSeries_step_down (a b : ℕ) {c d n : ℕ} {z : ℝ} apply le_sqrt_of_sq_le have hb' : 0 < (b : ℝ) := Nat.cast_pos.2 hb have hd' : 0 < (d : ℝ) := Nat.cast_pos.2 hd - rw [div_pow, add_div_eq_mul_add_div _ _ (ne_of_gt hd'), div_le_div_iff (pow_pos hb' _) hd'] + rw [div_pow, add_div_eq_mul_add_div _ _ (ne_of_gt hd'), div_le_div_iff₀ (pow_pos hb' _) hd'] exact mod_cast h section Tactic diff --git a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean index 64702155e2477..0b03b2f6c460d 100644 --- a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean +++ b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean @@ -600,7 +600,7 @@ theorem tendsto_translationNumber_of_dist_bounded_aux (x : ℕ → ℝ) (C : ℝ · exact fun n => C / 2 ^ n · intro n have : 0 < (2 ^ n : ℝ) := pow_pos zero_lt_two _ - convert (div_le_div_right this).2 (H (2 ^ n)) using 1 + convert (div_le_div_iff_of_pos_right this).2 (H (2 ^ n)) using 1 rw [transnumAuxSeq, Real.dist_eq, ← sub_div, abs_div, abs_of_pos this, Real.dist_eq] · exact mul_zero C ▸ tendsto_const_nhds.mul <| tendsto_inv_atTop_zero.comp <| tendsto_pow_atTop_atTop_of_one_lt one_lt_two @@ -673,7 +673,7 @@ theorem tendsto_translation_number₀' : dsimp have : (0 : ℝ) < n + 1 := n.cast_add_one_pos rw [Real.dist_eq, div_sub' _ _ _ (ne_of_gt this), abs_div, ← Real.dist_eq, abs_of_pos this, - Nat.cast_add_one, div_le_div_right this, ← Nat.cast_add_one] + Nat.cast_add_one, div_le_div_iff_of_pos_right this, ← Nat.cast_add_one] apply dist_pow_map_zero_mul_translationNumber_le theorem tendsto_translation_number₀ : Tendsto (fun n : ℕ => (f ^ n) 0 / n) atTop (𝓝 <| τ f) := diff --git a/Mathlib/Geometry/Euclidean/Inversion/Basic.lean b/Mathlib/Geometry/Euclidean/Inversion/Basic.lean index 4b75f996be818..ef7c492a1d056 100644 --- a/Mathlib/Geometry/Euclidean/Inversion/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Inversion/Basic.lean @@ -194,7 +194,7 @@ theorem mul_dist_le_mul_dist_add_mul_dist (a b c d : P) : rw [dist_inversion_inversion hb hd, dist_inversion_inversion hb hc, dist_inversion_inversion hc hd, one_pow] at H rw [← dist_pos] at hb hc hd - rw [← div_le_div_right (mul_pos hb (mul_pos hc hd))] + rw [← div_le_div_iff_of_pos_right (mul_pos hb (mul_pos hc hd))] convert H using 1 <;> (field_simp [hb.ne', hc.ne', hd.ne', dist_comm a]; ring) end EuclideanGeometry diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index 8142e71f6f653..f01189b606a68 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -6,6 +6,7 @@ Authors: Johannes Hölzl, Julian Kuelshammer import Mathlib.Algebra.CharP.Defs import Mathlib.Algebra.Group.Subgroup.Finite import Mathlib.Algebra.Order.Group.Action +import Mathlib.Algebra.Order.Ring.Abs import Mathlib.GroupTheory.Index import Mathlib.Order.Interval.Set.Infinite diff --git a/Mathlib/MeasureTheory/Integral/PeakFunction.lean b/Mathlib/MeasureTheory/Integral/PeakFunction.lean index 7d17886a1c7d2..9c33e5b413492 100644 --- a/Mathlib/MeasureTheory/Integral/PeakFunction.lean +++ b/Mathlib/MeasureTheory/Integral/PeakFunction.lean @@ -321,11 +321,11 @@ theorem tendsto_setIntegral_pow_smul_of_unique_maximum_of_isCompact_of_measure_n _ ≤ ∫ y in s, c y ^ n ∂μ := setIntegral_mono_set (I n) (J n) (Eventually.of_forall inter_subset_right) simp_rw [φ, ← div_eq_inv_mul, div_pow, div_div] - apply div_le_div (pow_nonneg t_pos n) _ _ B - · exact pow_le_pow_left₀ (hnc _ hx.1) (ht x hx) _ - · apply mul_pos (pow_pos (t_pos.trans_lt tt') _) (ENNReal.toReal_pos (hμ v v_open x₀_v).ne' _) - have : μ (v ∩ s) ≤ μ s := measure_mono inter_subset_right - exact ne_of_lt (lt_of_le_of_lt this hs.measure_lt_top) + have := ENNReal.toReal_pos (hμ v v_open x₀_v).ne' + ((measure_mono inter_subset_right).trans_lt hs.measure_lt_top).ne + gcongr + · exact hnc _ hx.1 + · exact ht x hx have N : Tendsto (fun n => (μ (v ∩ s)).toReal⁻¹ * (t / t') ^ n) atTop (𝓝 ((μ (v ∩ s)).toReal⁻¹ * 0)) := by diff --git a/Mathlib/MeasureTheory/Measure/Haar/Basic.lean b/Mathlib/MeasureTheory/Measure/Haar/Basic.lean index c09a144500014..546c35be2dda7 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Basic.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Basic.lean @@ -280,7 +280,7 @@ theorem prehaar_pos (K₀ : PositiveCompacts G) {U : Set G} (hU : (interior U).N theorem prehaar_mono {K₀ : PositiveCompacts G} {U : Set G} (hU : (interior U).Nonempty) {K₁ K₂ : Compacts G} (h : (K₁ : Set G) ⊆ K₂.1) : prehaar (K₀ : Set G) U K₁ ≤ prehaar (K₀ : Set G) U K₂ := by - simp only [prehaar]; rw [div_le_div_right] + simp only [prehaar]; rw [div_le_div_iff_of_pos_right] · exact mod_cast index_mono K₂.2 h hU · exact mod_cast index_pos K₀ hU @@ -293,7 +293,7 @@ theorem prehaar_self {K₀ : PositiveCompacts G} {U : Set G} (hU : (interior U). theorem prehaar_sup_le {K₀ : PositiveCompacts G} {U : Set G} (K₁ K₂ : Compacts G) (hU : (interior U).Nonempty) : prehaar (K₀ : Set G) U (K₁ ⊔ K₂) ≤ prehaar (K₀ : Set G) U K₁ + prehaar (K₀ : Set G) U K₂ := by - simp only [prehaar]; rw [div_add_div_same, div_le_div_right] + simp only [prehaar]; rw [div_add_div_same, div_le_div_iff_of_pos_right] · exact mod_cast index_union_le K₁ K₂ hU · exact mod_cast index_pos K₀ hU diff --git a/Mathlib/NumberTheory/DiophantineApproximation.lean b/Mathlib/NumberTheory/DiophantineApproximation.lean index 604980a89620e..cdccd001950cb 100644 --- a/Mathlib/NumberTheory/DiophantineApproximation.lean +++ b/Mathlib/NumberTheory/DiophantineApproximation.lean @@ -408,7 +408,7 @@ private theorem aux₁ : 0 < fract ξ := by have H : (2 * v - 1 : ℝ) < 1 := by refine (mul_lt_iff_lt_one_right hv₀).1 ((inv_lt_inv₀ hv₀ (mul_pos hv₁ hv₂)).1 (h.trans_le' ?_)) have h' : (⌊ξ⌋ : ℝ) - u / v = (⌊ξ⌋ * v - u) / v := by field_simp - rw [h', abs_div, abs_of_pos hv₀, ← one_div, div_le_div_right hv₀] + rw [h', abs_div, abs_of_pos hv₀, ← one_div, div_le_div_iff_of_pos_right hv₀] norm_cast rw [← zero_add (1 : ℤ), add_one_le_iff, abs_pos, sub_ne_zero] rintro rfl diff --git a/Mathlib/NumberTheory/FLT/Basic.lean b/Mathlib/NumberTheory/FLT/Basic.lean index 3630dbe1368d9..c6ebe0078285b 100644 --- a/Mathlib/NumberTheory/FLT/Basic.lean +++ b/Mathlib/NumberTheory/FLT/Basic.lean @@ -6,6 +6,7 @@ Authors: Kevin Buzzard, Yaël Dillies, Jineon Baek import Mathlib.Algebra.EuclideanDomain.Int import Mathlib.Algebra.GCDMonoid.Finset import Mathlib.Algebra.GCDMonoid.Nat +import Mathlib.Algebra.Order.Ring.Abs import Mathlib.RingTheory.PrincipalIdealDomain /-! diff --git a/Mathlib/NumberTheory/Harmonic/Defs.lean b/Mathlib/NumberTheory/Harmonic/Defs.lean index effef54ba3920..51c8f7036e358 100644 --- a/Mathlib/NumberTheory/Harmonic/Defs.lean +++ b/Mathlib/NumberTheory/Harmonic/Defs.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Koundinya Vajjha, Thomas Browning -/ import Mathlib.Algebra.BigOperators.Intervals +import Mathlib.Algebra.Order.Field.Basic import Mathlib.Tactic.Positivity.Finset /-! diff --git a/Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean b/Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean index 93d06720ea1b3..ae321a0cce2eb 100644 --- a/Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean +++ b/Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean @@ -274,7 +274,7 @@ lemma continuousOn_term_tsum : ContinuousOn term_tsum (Ici 1) := by refine setIntegral_mono_on ?_ ?_ measurableSet_Ioc (fun x hx ↦ ?_) · exact (term_welldef n.succ_pos (zero_lt_one.trans_le hs)).1 · exact (term_welldef n.succ_pos zero_lt_one).1 - · rw [div_le_div_left] -- leave side-goals to end and kill them all together + · rw [div_le_div_iff_of_pos_left] -- leave side-goals to end and kill them all together · apply rpow_le_rpow_of_exponent_le · exact (lt_of_le_of_lt (by simp) hx.1).le · linarith [mem_Ici.mp hs] diff --git a/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean b/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean index 5407eebea6353..81dee4c2c5ae9 100644 --- a/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean +++ b/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean @@ -519,7 +519,7 @@ lemma hasSum_int_completedCosZeta (a : ℝ) {s : ℂ} (hs : 1 < re s) : rw [mellin_div_const, completedCosZeta] congr 1 refine ((hurwitzEvenFEPair a).symm.hasMellin (?_ : 1 / 2 < (s / 2).re)).2.symm - rwa [div_ofNat_re, div_lt_div_right two_pos]] + rwa [div_ofNat_re, div_lt_div_iff_of_pos_right two_pos]] refine (hasSum_mellin_pi_mul_sq (zero_lt_one.trans hs) hF ?_).congr_fun fun n ↦ ?_ · apply (((summable_one_div_int_add_rpow 0 s.re).mpr hs).div_const 2).of_norm_bounded intro i @@ -560,7 +560,7 @@ lemma hasSum_int_completedHurwitzZetaEven (a : ℝ) {s : ℂ} (hs : 1 < re s) : ↑(if (a : UnitAddCircle) = 0 then 1 else 0 : ℝ)) / 2) (s / 2) by simp_rw [mellin_div_const, apply_ite ofReal, ofReal_one, ofReal_zero] refine congr_arg (· / 2) ((hurwitzEvenFEPair a).hasMellin (?_ : 1 / 2 < (s / 2).re)).2.symm - rwa [div_ofNat_re, div_lt_div_right two_pos]] + rwa [div_ofNat_re, div_lt_div_iff_of_pos_right two_pos]] refine (hasSum_mellin_pi_mul_sq (zero_lt_one.trans hs) hF ?_).congr_fun fun n ↦ ?_ · simp_rw [← mul_one_div ‖_‖] apply Summable.mul_left diff --git a/Mathlib/NumberTheory/Modular.lean b/Mathlib/NumberTheory/Modular.lean index 6426b4dac0aad..8f86113a1bed0 100644 --- a/Mathlib/NumberTheory/Modular.lean +++ b/Mathlib/NumberTheory/Modular.lean @@ -279,7 +279,7 @@ theorem exists_max_im : ∃ g : SL(2, ℤ), ∀ g' : SL(2, ℤ), (g' • z).im obtain ⟨g, -, hg⟩ := bottom_row_surj hp_coprime refine ⟨g, fun g' => ?_⟩ rw [ModularGroup.im_smul_eq_div_normSq, ModularGroup.im_smul_eq_div_normSq, - div_le_div_left] + div_le_div_iff_of_pos_left] · simpa [← hg] using hp (g' 1) (bottom_row_coprime g') · exact z.im_pos · exact normSq_denom_pos g' z diff --git a/Mathlib/NumberTheory/Padics/Hensel.lean b/Mathlib/NumberTheory/Padics/Hensel.lean index e45dc9dad9d42..74e302eb5dbe2 100644 --- a/Mathlib/NumberTheory/Padics/Hensel.lean +++ b/Mathlib/NumberTheory/Padics/Hensel.lean @@ -267,7 +267,7 @@ private theorem newton_seq_succ_dist (n : ℕ) : newton_seq_norm_eq hnorm _ _ = ‖F.eval (newton_seq n)‖ / ‖F.derivative.eval a‖ := by rw [newton_seq_deriv_norm] _ ≤ ‖F.derivative.eval a‖ ^ 2 * T ^ 2 ^ n / ‖F.derivative.eval a‖ := - ((div_le_div_right (deriv_norm_pos hnorm)).2 (newton_seq_norm_le hnorm _)) + ((div_le_div_iff_of_pos_right (deriv_norm_pos hnorm)).2 (newton_seq_norm_le hnorm _)) _ = ‖F.derivative.eval a‖ * T ^ 2 ^ n := div_sq_cancel _ _ private theorem newton_seq_dist_aux (n : ℕ) : diff --git a/Mathlib/NumberTheory/Pell.lean b/Mathlib/NumberTheory/Pell.lean index e83f9f210649a..c3559d516dd99 100644 --- a/Mathlib/NumberTheory/Pell.lean +++ b/Mathlib/NumberTheory/Pell.lean @@ -337,7 +337,8 @@ theorem exists_of_not_isSquare (h₀ : 0 < d) (hd : ¬IsSquare d) : refine Infinite.mono (fun q h => ?_) (infinite_rat_abs_sub_lt_one_div_den_sq_of_irrational hξ) have h0 : 0 < (q.2 : ℝ) ^ 2 := pow_pos (Nat.cast_pos.mpr q.pos) 2 have h1 : (q.num : ℝ) / (q.den : ℝ) = q := mod_cast q.num_div_den - rw [mem_setOf, abs_sub_comm, ← @Int.cast_lt ℝ, ← div_lt_div_right (abs_pos_of_pos h0)] + rw [mem_setOf, abs_sub_comm, ← @Int.cast_lt ℝ, + ← div_lt_div_iff_of_pos_right (abs_pos_of_pos h0)] push_cast rw [← abs_div, abs_sq, sub_div, mul_div_cancel_right₀ _ h0.ne', ← div_pow, h1, ← sq_sqrt (Int.cast_pos.mpr h₀).le, sq_sub_sq, abs_mul, ← mul_one_div] diff --git a/Mathlib/NumberTheory/Transcendental/Liouville/Basic.lean b/Mathlib/NumberTheory/Transcendental/Liouville/Basic.lean index 04403810047a7..b3587eda5cecf 100644 --- a/Mathlib/NumberTheory/Transcendental/Liouville/Basic.lean +++ b/Mathlib/NumberTheory/Transcendental/Liouville/Basic.lean @@ -47,7 +47,7 @@ protected theorem irrational {x : ℝ} (h : Liouville x) : Irrational x := by have bq0 : (0 : ℝ) < b * q := mul_pos (Nat.cast_pos.mpr bN0.bot_lt) qR0 -- At a1, clear denominators... replace a1 : |a * q - b * p| * q ^ (b + 1) < b * q := by - rw [div_sub_div _ _ b0 qR0.ne', abs_div, div_lt_div_iff (abs_pos.mpr bq0.ne') (pow_pos qR0 _), + rw [div_sub_div _ _ b0 qR0.ne', abs_div, div_lt_div_iff₀ (abs_pos.mpr bq0.ne') (pow_pos qR0 _), abs_of_pos bq0, one_mul] at a1 exact mod_cast a1 -- At a0, clear denominators... diff --git a/Mathlib/NumberTheory/Transcendental/Liouville/LiouvilleNumber.lean b/Mathlib/NumberTheory/Transcendental/Liouville/LiouvilleNumber.lean index 72b1c28fd82b7..2b9b52c60e496 100644 --- a/Mathlib/NumberTheory/Transcendental/Liouville/LiouvilleNumber.lean +++ b/Mathlib/NumberTheory/Transcendental/Liouville/LiouvilleNumber.lean @@ -136,7 +136,7 @@ theorem aux_calc (n : ℕ) {m : ℝ} (hm : 2 ≤ m) : -- [NB: in this block, I do not follow the brace convention for subgoals -- I wait until -- I solve all extraneous goals at once with `exact pow_pos (zero_lt_two.trans_le hm) _`.] -- Clear denominators and massage* - apply (div_le_div_iff _ _).mpr + apply (div_le_div_iff₀ _ _).mpr focus conv_rhs => rw [one_mul, mul_add, pow_add, mul_one, pow_mul, mul_comm, ← pow_mul] -- the second factors coincide, so we prove the inequality of the first factors* diff --git a/Mathlib/NumberTheory/Transcendental/Liouville/Measure.lean b/Mathlib/NumberTheory/Transcendental/Liouville/Measure.lean index ca97dfa42b9e9..8421fa4e65ef5 100644 --- a/Mathlib/NumberTheory/Transcendental/Liouville/Measure.lean +++ b/Mathlib/NumberTheory/Transcendental/Liouville/Measure.lean @@ -58,8 +58,8 @@ theorem setOf_liouvilleWith_subset_aux : _ ≤ (b : ℝ) ^ (2 + 1 / (n + 1 : ℕ) : ℝ) := rpow_le_rpow_of_exponent_le hb (one_le_two.trans ?_) simpa using n.cast_add_one_pos.le - rw [sub_div' _ _ _ hb0.ne', abs_div, abs_of_pos hb0, div_lt_div_right hb0, abs_sub_lt_iff, - sub_lt_iff_lt_add, sub_lt_iff_lt_add, ← sub_lt_iff_lt_add'] at hlt + rw [sub_div' _ _ _ hb0.ne', abs_div, abs_of_pos hb0, div_lt_div_iff_of_pos_right hb0, + abs_sub_lt_iff, sub_lt_iff_lt_add, sub_lt_iff_lt_add, ← sub_lt_iff_lt_add'] at hlt rw [Finset.mem_Icc, ← Int.lt_add_one_iff, ← Int.lt_add_one_iff, ← neg_lt_iff_pos_add, add_comm, ← @Int.cast_lt ℝ, ← @Int.cast_lt ℝ] push_cast diff --git a/Mathlib/Order/Interval/Set/IsoIoo.lean b/Mathlib/Order/Interval/Set/IsoIoo.lean index f44e1e4822571..649d99a424435 100644 --- a/Mathlib/Order/Interval/Set/IsoIoo.lean +++ b/Mathlib/Order/Interval/Set/IsoIoo.lean @@ -35,7 +35,7 @@ def orderIsoIooNegOneOne (k : Type*) [LinearOrderedField k] : k ≃o Ioo (-1 : k · intro x simp only [abs_neg, neg_div] · rintro x (hx : 0 ≤ x) y (hy : 0 ≤ y) hxy - simp [abs_of_nonneg, mul_add, mul_comm x y, div_lt_div_iff, hx.trans_lt (lt_one_add _), + simp [abs_of_nonneg, mul_add, mul_comm x y, div_lt_div_iff₀, hx.trans_lt (lt_one_add _), hy.trans_lt (lt_one_add _), *] · refine fun x ↦ Subtype.ext ?_ have : 0 < 1 - |(x : k)| := sub_pos.2 (abs_lt.2 x.2) diff --git a/Mathlib/RingTheory/Multiplicity.lean b/Mathlib/RingTheory/Multiplicity.lean index 85ab17ad0be89..c095227abd535 100644 --- a/Mathlib/RingTheory/Multiplicity.lean +++ b/Mathlib/RingTheory/Multiplicity.lean @@ -5,8 +5,9 @@ Authors: Robert Y. Lewis, Chris Hughes, Daniel Weber -/ import Mathlib.Algebra.Associated.Basic import Mathlib.Algebra.BigOperators.Group.Finset -import Mathlib.Tactic.Linarith +import Mathlib.Algebra.Ring.Divisibility.Basic import Mathlib.Data.ENat.Basic +import Mathlib.Tactic.Linarith /-! # Multiplicity of a divisor diff --git a/Mathlib/RingTheory/Polynomial/Hermite/Basic.lean b/Mathlib/RingTheory/Polynomial/Hermite/Basic.lean index fa81775947129..13f33a3208377 100644 --- a/Mathlib/RingTheory/Polynomial/Hermite/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Hermite/Basic.lean @@ -3,7 +3,6 @@ Copyright (c) 2023 Luke Mantle. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Luke Mantle -/ -import Mathlib.Algebra.Order.Ring.Abs import Mathlib.Algebra.Polynomial.Derivative import Mathlib.Data.Nat.Factorial.DoubleFactorial diff --git a/Mathlib/RingTheory/PowerSeries/WellKnown.lean b/Mathlib/RingTheory/PowerSeries/WellKnown.lean index 292659bb8862f..ef67977a0a5d2 100644 --- a/Mathlib/RingTheory/PowerSeries/WellKnown.lean +++ b/Mathlib/RingTheory/PowerSeries/WellKnown.lean @@ -5,7 +5,6 @@ Authors: Yury Kudryashov -/ import Mathlib.Algebra.Algebra.Rat import Mathlib.Algebra.BigOperators.NatAntidiagonal -import Mathlib.Algebra.Order.Ring.Abs import Mathlib.Data.Nat.Choose.Sum import Mathlib.RingTheory.PowerSeries.Basic diff --git a/Mathlib/SetTheory/Ordinal/Notation.lean b/Mathlib/SetTheory/Ordinal/Notation.lean index f5a5db19acb40..02806888024eb 100644 --- a/Mathlib/SetTheory/Ordinal/Notation.lean +++ b/Mathlib/SetTheory/Ordinal/Notation.lean @@ -3,6 +3,7 @@ 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.Algebra.Ring.Divisibility.Basic import Mathlib.Data.Ordering.Lemmas import Mathlib.Data.PNat.Basic import Mathlib.SetTheory.Ordinal.Principal diff --git a/Mathlib/SetTheory/Surreal/Dyadic.lean b/Mathlib/SetTheory/Surreal/Dyadic.lean index 66e64a35705b5..b1dc9275095fb 100644 --- a/Mathlib/SetTheory/Surreal/Dyadic.lean +++ b/Mathlib/SetTheory/Surreal/Dyadic.lean @@ -5,12 +5,12 @@ Authors: Apurva Nakade -/ import Mathlib.Algebra.Algebra.Defs import Mathlib.Algebra.Order.Group.Basic +import Mathlib.Algebra.Ring.Regular import Mathlib.GroupTheory.MonoidLocalization.Away import Mathlib.RingTheory.Localization.Defs import Mathlib.SetTheory.Game.Birthday import Mathlib.SetTheory.Surreal.Multiplication import Mathlib.Tactic.Linarith -import Mathlib.Algebra.Ring.Regular /-! # Dyadic numbers diff --git a/Mathlib/Tactic/LinearCombination.lean b/Mathlib/Tactic/LinearCombination.lean index f1c65e78a805b..5c891e021eae9 100644 --- a/Mathlib/Tactic/LinearCombination.lean +++ b/Mathlib/Tactic/LinearCombination.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Abby J. Goldberg, Mario Carneiro, Heather Macbeth -/ import Mathlib.Tactic.LinearCombination.Lemmas +import Mathlib.Tactic.Positivity.Core import Mathlib.Tactic.Ring import Mathlib.Tactic.Ring.Compare diff --git a/Mathlib/Tactic/LinearCombination/Lemmas.lean b/Mathlib/Tactic/LinearCombination/Lemmas.lean index 2d93dbab83217..e22bec542ae8b 100644 --- a/Mathlib/Tactic/LinearCombination/Lemmas.lean +++ b/Mathlib/Tactic/LinearCombination/Lemmas.lean @@ -3,8 +3,8 @@ Copyright (c) 2022 Abby J. Goldberg. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Abby J. Goldberg, Mario Carneiro, Heather Macbeth -/ +import Mathlib.Algebra.Order.Field.Defs import Mathlib.Data.Ineq -import Mathlib.Algebra.Order.Field.Basic /-! # Lemmas for the linear_combination tactic diff --git a/Mathlib/Tactic/NormNum/GCD.lean b/Mathlib/Tactic/NormNum/GCD.lean index 7594701886b26..cf43344d0f6e3 100644 --- a/Mathlib/Tactic/NormNum/GCD.lean +++ b/Mathlib/Tactic/NormNum/GCD.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Kyle Miller, Eric Wieser -/ +import Mathlib.Algebra.Ring.Divisibility.Basic import Mathlib.Data.Int.GCD import Mathlib.Tactic.NormNum diff --git a/Mathlib/Topology/Connected/PathConnected.lean b/Mathlib/Topology/Connected/PathConnected.lean index 82ff89fe8d4c4..c5d94601f96f7 100644 --- a/Mathlib/Topology/Connected/PathConnected.lean +++ b/Mathlib/Topology/Connected/PathConnected.lean @@ -330,7 +330,7 @@ theorem trans_range {a b c : X} (γ₁ : Path a b) (γ₂ : Path b c) : rwa [coe_mk_mk, Function.comp_apply, if_neg h] at hxt · rintro x (⟨⟨t, ht0, ht1⟩, hxt⟩ | ⟨⟨t, ht0, ht1⟩, hxt⟩) · use ⟨t / 2, ⟨by linarith, by linarith⟩⟩ - have : t / 2 ≤ 1 / 2 := (div_le_div_right (zero_lt_two : (0 : ℝ) < 2)).mpr ht1 + have : t / 2 ≤ 1 / 2 := (div_le_div_iff_of_pos_right (zero_lt_two : (0 : ℝ) < 2)).mpr ht1 rw [coe_mk_mk, Function.comp_apply, if_pos this, Subtype.coe_mk] ring_nf rwa [γ₁.extend_extends] diff --git a/Mathlib/Topology/MetricSpace/Contracting.lean b/Mathlib/Topology/MetricSpace/Contracting.lean index e70e951d65294..a4c5c2b74235d 100644 --- a/Mathlib/Topology/MetricSpace/Contracting.lean +++ b/Mathlib/Topology/MetricSpace/Contracting.lean @@ -266,7 +266,7 @@ theorem dist_fixedPoint_fixedPoint_of_dist_le' (g : α → α) {x y} (hx : IsFix dist x y = dist y x := dist_comm x y _ ≤ dist y (f y) / (1 - K) := hf.dist_le_of_fixedPoint y hx _ = dist (f y) (g y) / (1 - K) := by rw [hy.eq, dist_comm] - _ ≤ C / (1 - K) := (div_le_div_right hf.one_sub_K_pos).2 (hfg y) + _ ≤ C / (1 - K) := (div_le_div_iff_of_pos_right hf.one_sub_K_pos).2 (hfg y) variable [Nonempty α] [CompleteSpace α] variable (f) diff --git a/Mathlib/Topology/TietzeExtension.lean b/Mathlib/Topology/TietzeExtension.lean index 37c54d6ae46ae..84b683fbeb2fe 100644 --- a/Mathlib/Topology/TietzeExtension.lean +++ b/Mathlib/Topology/TietzeExtension.lean @@ -181,7 +181,7 @@ theorem tietze_extension_step (f : X →ᵇ ℝ) (e : C(X, Y)) (he : IsClosedEmb are disjoint, hence by Urysohn's lemma there exists a function `g` that is equal to `-‖f‖ / 3` on the former set and is equal to `‖f‖ / 3` on the latter set. This function `g` satisfies the assertions of the lemma. -/ - have hf3 : -‖f‖ / 3 < ‖f‖ / 3 := (div_lt_div_right h3).2 (Left.neg_lt_self hf) + have hf3 : -‖f‖ / 3 < ‖f‖ / 3 := (div_lt_div_iff_of_pos_right h3).2 (Left.neg_lt_self hf) have hc₁ : IsClosed (e '' (f ⁻¹' Iic (-‖f‖ / 3))) := he.isClosedMap _ (isClosed_Iic.preimage f.continuous) have hc₂ : IsClosed (e '' (f ⁻¹' Ici (‖f‖ / 3))) := diff --git a/MathlibTest/apply_rules.lean b/MathlibTest/apply_rules.lean index d621b7bb85ae8..ab93b4c736a69 100644 --- a/MathlibTest/apply_rules.lean +++ b/MathlibTest/apply_rules.lean @@ -49,4 +49,4 @@ example [LinearOrderedField α] {a b : α} (hb : 0 ≤ b) (hab : a ≤ b) : a / fail_if_success apply_rules (config := { transparency := .reducible }) [mul_le_mul] guard_target = a / 2 ≤ b / 2 - exact div_le_div hb hab zero_lt_two le_rfl + exact div_le_div₀ hb hab zero_lt_two le_rfl diff --git a/scripts/no_lints_prime_decls.txt b/scripts/no_lints_prime_decls.txt index f00efff46d47e..7a7e8fb75bfe5 100644 --- a/scripts/no_lints_prime_decls.txt +++ b/scripts/no_lints_prime_decls.txt @@ -1088,7 +1088,6 @@ div_eq_iff_eq_mul' div_eq_of_eq_mul' div_eq_of_eq_mul'' div_le_div'' -div_le_div_iff' div_le_div_left' div_le_div_right' div_left_inj' @@ -1097,6 +1096,7 @@ div_le_iff_le_mul' div_le_iff_of_neg' div_le_one' div_lt_div' +div_lt_div₀' div_lt_div'' div_lt_div_iff' div_lt_div_left' diff --git a/scripts/noshake.json b/scripts/noshake.json index a1d6020a42bd4..736e94310eaf9 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -340,6 +340,7 @@ "Mathlib.Order.BoundedOrder": ["Mathlib.Tactic.Finiteness.Attr"], "Mathlib.Order.Basic": ["Batteries.Tactic.Classical", "Mathlib.Tactic.GCongr.Core"], + "Mathlib.NumberTheory.Harmonic.Defs": ["Mathlib.Algebra.Order.Field.Basic"], "Mathlib.NumberTheory.Cyclotomic.Basic": ["Mathlib.Init.Core"], "Mathlib.NumberTheory.ArithmeticFunction": ["Mathlib.Tactic.ArithMult"], "Mathlib.MeasureTheory.MeasurableSpace.Defs": ["Mathlib.Tactic.FunProp.Attr"], @@ -392,6 +393,8 @@ "Mathlib.Control.Monad.Cont": ["Batteries.Lean.Except", "Batteries.Tactic.Congr"], "Mathlib.Control.Basic": ["Mathlib.Logic.Function.Defs"], + "Mathlib.Combinatorics.SimpleGraph.Density": + ["Mathlib.Algebra.Order.Field.Basic"], "Mathlib.Combinatorics.SetFamily.AhlswedeZhang": ["Mathlib.Algebra.Order.Field.Basic"], "Mathlib.CategoryTheory.Sites.IsSheafFor": From 4e69bf9a2bbce8bca6666996a53d842121ddf304 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 19 Nov 2024 22:02:12 +0000 Subject: [PATCH 125/829] feat(SetTheory/ZFC/Ordinal): Initial development of ordinals in ZFSet (#15793) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This is ported from the Mathlib 3 branch [`von_neumann_v2`](https://github.com/leanprover-community/mathlib3/blob/5b3c1d3d66838f426b44cae25b615f26438c8acb/src/set_theory/zfc/ordinal.lean). This is meant as a small, initial PR to iron out any kinks (with the definition, notation, auto-ported lemmas, etc.) before fully developing the file. Of particular note is the quite bare-bones definition of an ordinal. Surprisingly, our weakened transitivity assumption is equivalent to the much stronger claim that ordinals are well-ordered under `∈`; see #17001 for a proof. --- Mathlib/SetTheory/ZFC/Ordinal.lean | 68 ++++++++++++++++++++++++------ 1 file changed, 54 insertions(+), 14 deletions(-) diff --git a/Mathlib/SetTheory/ZFC/Ordinal.lean b/Mathlib/SetTheory/ZFC/Ordinal.lean index 56dbcfb274d7a..bd816b21e5dac 100644 --- a/Mathlib/SetTheory/ZFC/Ordinal.lean +++ b/Mathlib/SetTheory/ZFC/Ordinal.lean @@ -3,26 +3,27 @@ Copyright (c) 2022 Violeta Hernández Palacios. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Violeta Hernández Palacios -/ +import Mathlib.Order.RelIso.Set import Mathlib.SetTheory.ZFC.Basic /-! # Von Neumann ordinals This file works towards the development of von Neumann ordinals, i.e. transitive sets, well-ordered -under `∈`. We currently only have an initial development of transitive sets. +under `∈`. -Further development can be found on the branch `von_neumann_v2`. +We currently only have an initial development of transitive sets and ordinals. Further development +can be found on the Mathlib 3 branch `von_neumann_v2`. ## Definitions - `ZFSet.IsTransitive` means that every element of a set is a subset. +- `ZFSet.IsOrdinal` means that the set is transitive and well-ordered under `∈`. ## TODO -- Define von Neumann ordinals. -- Define the basic arithmetic operations on ordinals from a purely set-theoretic perspective. -- Prove the equivalences between these definitions and those provided in - `SetTheory/Ordinal/Arithmetic.lean`. +- Define von Neumann ordinals and the von Neumann hierarchy. +- Build correspondences between these notions and those of the standard `Ordinal` type. -/ @@ -38,10 +39,12 @@ def IsTransitive (x : ZFSet) : Prop := ∀ y ∈ x, y ⊆ x @[simp] -theorem empty_isTransitive : IsTransitive ∅ := fun y hy => (not_mem_empty y hy).elim +theorem isTransitive_empty : IsTransitive ∅ := fun y hy => (not_mem_empty y hy).elim -theorem IsTransitive.subset_of_mem (h : x.IsTransitive) : y ∈ x → y ⊆ x := - h y +@[deprecated isTransitive_empty (since := "2024-09-21")] +alias empty_isTransitive := isTransitive_empty + +theorem IsTransitive.subset_of_mem (h : x.IsTransitive) : y ∈ x → y ⊆ x := h y theorem isTransitive_iff_mem_trans : z.IsTransitive ↔ ∀ {x y : ZFSet}, x ∈ y → y ∈ z → x ∈ z := ⟨fun h _ _ hx hy => h.subset_of_mem hy hx, fun H _ hx _ hy => H hy hx⟩ @@ -66,7 +69,7 @@ theorem IsTransitive.sUnion' (H : ∀ y ∈ x, IsTransitive y) : protected theorem IsTransitive.union (hx : x.IsTransitive) (hy : y.IsTransitive) : (x ∪ y).IsTransitive := by rw [← sUnion_pair] - apply IsTransitive.sUnion' fun z => _ + apply IsTransitive.sUnion' intro rw [mem_pair] rintro (rfl | rfl) @@ -77,10 +80,12 @@ protected theorem IsTransitive.powerset (h : x.IsTransitive) : (powerset x).IsTr rw [mem_powerset] at hy ⊢ exact h.subset_of_mem (hy hz) -theorem isTransitive_iff_sUnion_subset : x.IsTransitive ↔ (⋃₀ x : ZFSet) ⊆ x := - ⟨fun h y hy => by - rcases mem_sUnion.1 hy with ⟨z, hz, hz'⟩ - exact h.mem_trans hz' hz, fun H _ hy _ hz => H <| mem_sUnion_of_mem hz hy⟩ +theorem isTransitive_iff_sUnion_subset : x.IsTransitive ↔ (⋃₀ x : ZFSet) ⊆ x := by + constructor <;> + intro h y hy + · obtain ⟨z, hz, hz'⟩ := mem_sUnion.1 hy + exact h.mem_trans hz' hz + · exact fun z hz ↦ h <| mem_sUnion_of_mem hz hy alias ⟨IsTransitive.sUnion_subset, _⟩ := isTransitive_iff_sUnion_subset @@ -89,4 +94,39 @@ theorem isTransitive_iff_subset_powerset : x.IsTransitive ↔ x ⊆ powerset x : alias ⟨IsTransitive.subset_powerset, _⟩ := isTransitive_iff_subset_powerset +/-- A set `x` is a von Neumann ordinal when it's a transitive set, that's transitive under `∈`. We +will prove that this further implies that `x` is well-ordered under `∈`. + +The transitivity condition `a ∈ b → b ∈ c → a ∈ c` can be written without assuming `a ∈ x` and +`b ∈ x`. The lemma `isOrdinal_iff_isTrans` shows this condition is equivalent to the usual one. -/ +structure IsOrdinal (x : ZFSet) : Prop where + /-- An ordinal is a transitive set. -/ + isTransitive : x.IsTransitive + /-- The membership operation within an ordinal is transitive. -/ + mem_trans' {y z w : ZFSet} : y ∈ z → z ∈ w → w ∈ x → y ∈ w + +namespace IsOrdinal + +theorem subset_of_mem (h : x.IsOrdinal) : y ∈ x → y ⊆ x := + h.isTransitive.subset_of_mem + +theorem mem_trans (h : z.IsOrdinal) : x ∈ y → y ∈ z → x ∈ z := + h.isTransitive.mem_trans + +protected theorem isTrans (h : x.IsOrdinal) : IsTrans x.toSet (Subrel (· ∈ ·) _) := + ⟨fun _ _ c hab hbc => h.mem_trans' hab hbc c.2⟩ + +end IsOrdinal + +/-- The simplified form of transitivity used within `IsOrdinal` yields an equivalent definition to +the standard one. -/ +theorem isOrdinal_iff_isTrans : + x.IsOrdinal ↔ x.IsTransitive ∧ IsTrans x.toSet (Subrel (· ∈ ·) _) := by + use fun h => ⟨h.isTransitive, h.isTrans⟩ + rintro ⟨h₁, ⟨h₂⟩⟩ + use h₁ + intro y z w hyz hzw hwx + have hzx := h₁.mem_trans hzw hwx + exact h₂ ⟨y, h₁.mem_trans hyz hzx⟩ ⟨z, hzx⟩ ⟨w, hwx⟩ hyz hzw + end ZFSet From c3a2d7bbf53eb210617d830940228487ff55d3d6 Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 19 Nov 2024 22:30:50 +0000 Subject: [PATCH 126/829] feat: `#parse` -- a command to parse text and log outputs (#16305) This is a spin-off of #16230. It introduces a function written by Matt that allows to use `#guard_msgs` to check parsing errors. The new command is called `#guard_exceptions` and mimics very closely `#guard_msgs`. As a way of showing how to use it, the PR also adds tests for Stacks Project Tags. Co-authored-by: Eric Wieser Co-authored-by: Matthew Ballard --- Mathlib.lean | 1 + Mathlib/Util/ParseCommand.lean | 57 ++++++++++++++++++++++++++++++++ MathlibTest/StacksAttribute.lean | 16 +++++++++ 3 files changed, 74 insertions(+) create mode 100644 Mathlib/Util/ParseCommand.lean diff --git a/Mathlib.lean b/Mathlib.lean index 7684aafa754e8..3ca25acb72351 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -5277,6 +5277,7 @@ import Mathlib.Util.IncludeStr import Mathlib.Util.LongNames import Mathlib.Util.MemoFix import Mathlib.Util.Notation3 +import Mathlib.Util.ParseCommand import Mathlib.Util.Qq import Mathlib.Util.SleepHeartbeats import Mathlib.Util.Superscript diff --git a/Mathlib/Util/ParseCommand.lean b/Mathlib/Util/ParseCommand.lean new file mode 100644 index 0000000000000..d3c71c46e9aac --- /dev/null +++ b/Mathlib/Util/ParseCommand.lean @@ -0,0 +1,57 @@ +/- +Copyright (c) 2024 Matthew Robert Ballard. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Matthew Robert Ballard, Damiano Testa +-/ + +import Lean.Elab.Command +import Mathlib.Init + +/-! +# `#parse` -- a command to parse text and log outputs +-/ + +namespace Mathlib.GuardExceptions + +open Lean Parser Elab Command +/-- `captureException env s input` uses the given `Environment` `env` to parse the `String` `input` +using the `ParserFn` `s`. + +This is a variation of `Lean.Parser.runParserCategory`. +-/ +def captureException (env : Environment) (s : ParserFn) (input : String) : Except String Syntax := + let ictx := mkInputContext input "" + let s := s.run ictx { env, options := {} } (getTokenTable env) (mkParserState input) + if !s.allErrors.isEmpty then + .error (s.toErrorMsg ictx) + else if ictx.input.atEnd s.pos then + .ok s.stxStack.back + else + .error ((s.mkError "end of input").toErrorMsg ictx) + +/-- `#parse parserFnId => str` allows to capture parsing exceptions. +`parserFnId` is the identifier of a `ParserFn` and `str` is the string that +`parserFnId` should parse. + +If the parse is successful, then the output is logged; +if the parse is successful, then the output is captured in an exception. + +In either case, `#guard_msgs` can then be used to capture the resulting parsing errors. + +For instance, `#parse` can be used as follows +```lean +/-- error: :1:3: Stacks tags must be exactly 4 characters -/ +#guard_msgs in #parse Mathlib.Stacks.stacksTagFn => "A05" +``` +-/ +syntax (name := parseCmd) "#parse " ident " => " str : command + +@[inherit_doc parseCmd] +elab_rules : command + | `(command| #parse $parserFnId => $str) => do + elabCommand <| ← `(command| + run_cmd do + let exc ← Lean.ofExcept <| captureException (← getEnv) $parserFnId $str + logInfo $str) + +end Mathlib.GuardExceptions diff --git a/MathlibTest/StacksAttribute.lean b/MathlibTest/StacksAttribute.lean index 3b70bdb597610..6fbee7f267bc2 100644 --- a/MathlibTest/StacksAttribute.lean +++ b/MathlibTest/StacksAttribute.lean @@ -1,4 +1,5 @@ import Mathlib.Tactic.StacksAttribute +import Mathlib.Util.ParseCommand /-- info: No tags found. -/ #guard_msgs in @@ -30,6 +31,21 @@ example : True := .intro @[stacks 0X14 "I can also have a comment"] example : True := .intro +@[stacks 0BR2, stacks 0X14 "I can also have a comment"] +example : True := .intro + +@[stacks 0X14 "I can also have a comment"] +example : True := .intro + +/-- error: :1:3: Stacks tags must be exactly 4 characters -/ +#guard_msgs in #parse Mathlib.StacksTag.stacksTagFn => "A05" + +/-- error: :1:4: Stacks tags must consist only of digits and uppercase letters. -/ +#guard_msgs in #parse Mathlib.StacksTag.stacksTagFn => "A05b" + +/-- info: 0BD5 -/ +#guard_msgs in #parse Mathlib.StacksTag.stacksTagFn => "0BD5" + /-- info: [Stacks Tag A04Q](https://stacks.math.columbia.edu/tag/A04Q) corresponds to declaration 'X.tagged'. (A comment) From 25b91b6001b15abdba13e10b8dcf838cfe9d292f Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Wed, 20 Nov 2024 00:01:29 +0000 Subject: [PATCH 127/829] feat: commutativity of `NonUnital{Star}Algebra.adjoin` (#18612) --- .../Algebra/Algebra/NonUnitalSubalgebra.lean | 51 +++++++++++++ Mathlib/Algebra/Star/NonUnitalSubalgebra.lean | 72 +++++++++++++++++++ 2 files changed, 123 insertions(+) diff --git a/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean b/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean index af1d1bc38d9af..bea24677990a7 100644 --- a/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean +++ b/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean @@ -1131,6 +1131,57 @@ end Centralizer end NonUnitalSubalgebra +namespace NonUnitalAlgebra + +open NonUnitalSubalgebra + +variable {R A : Type*} [CommSemiring R] [NonUnitalSemiring A] +variable [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] + +variable (R) in +lemma adjoin_le_centralizer_centralizer (s : Set A) : + adjoin R s ≤ centralizer R (centralizer R s) := + adjoin_le Set.subset_centralizer_centralizer + +lemma commute_of_mem_adjoin_of_forall_mem_commute {a b : A} {s : Set A} + (hb : b ∈ adjoin R s) (h : ∀ b ∈ s, Commute a b) : + Commute a b := by + have : a ∈ centralizer R s := by simpa only [Commute.symm_iff (a := a)] using h + exact adjoin_le_centralizer_centralizer R s hb a this + +lemma commute_of_mem_adjoin_singleton_of_commute {a b c : A} + (hc : c ∈ adjoin R {b}) (h : Commute a b) : + Commute a c := + commute_of_mem_adjoin_of_forall_mem_commute hc <| by simpa + +lemma commute_of_mem_adjoin_self {a b : A} (hb : b ∈ adjoin R {a}) : + Commute a b := + commute_of_mem_adjoin_singleton_of_commute hb rfl + +variable (R) in + +/-- If all elements of `s : Set A` commute pairwise, then `adjoin R s` is a non-unital commutative +semiring. + +See note [reducible non-instances]. -/ +abbrev adjoinNonUnitalCommSemiringOfComm {s : Set A} (hcomm : ∀ a ∈ s, ∀ b ∈ s, a * b = b * a) : + NonUnitalCommSemiring (adjoin R s) := + { (adjoin R s).toNonUnitalSemiring with + mul_comm := fun ⟨_, h₁⟩ ⟨_, h₂⟩ ↦ + have := adjoin_le_centralizer_centralizer R s + Subtype.ext <| Set.centralizer_centralizer_comm_of_comm hcomm _ (this h₁) _ (this h₂) } + +/-- If all elements of `s : Set A` commute pairwise, then `adjoin R s` is a non-unital commutative +ring. + +See note [reducible non-instances]. -/ +abbrev adjoinNonUnitalCommRingOfComm (R : Type*) {A : Type*} [CommRing R] [NonUnitalRing A] + [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] {s : Set A} + (hcomm : ∀ a ∈ s, ∀ b ∈ s, a * b = b * a) : NonUnitalCommRing (adjoin R s) := + { (adjoin R s).toNonUnitalRing, adjoinNonUnitalCommSemiringOfComm R hcomm with } + +end NonUnitalAlgebra + section Nat variable {R : Type*} [NonUnitalNonAssocSemiring R] diff --git a/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean b/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean index 86f059f7cf0fe..0f0476488ad54 100644 --- a/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean +++ b/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean @@ -6,6 +6,7 @@ Authors: Jireh Loreaux import Mathlib.Algebra.Algebra.NonUnitalSubalgebra import Mathlib.Algebra.Star.StarAlgHom import Mathlib.Algebra.Star.Center +import Mathlib.Algebra.Star.SelfAdjoint /-! # Non-unital Star Subalgebras @@ -1130,6 +1131,77 @@ theorem centralizer_le (s t : Set A) (h : s ⊆ t) : centralizer R t ≤ central theorem centralizer_univ : centralizer R Set.univ = center R A := SetLike.ext' <| by rw [coe_centralizer, Set.univ_union, coe_center, Set.centralizer_univ] +theorem centralizer_toNonUnitalSubalgebra (s : Set A) : + (centralizer R s).toNonUnitalSubalgebra = NonUnitalSubalgebra.centralizer R (s ∪ star s):= + rfl + +theorem coe_centralizer_centralizer (s : Set A) : + (centralizer R (centralizer R s : Set A)) = (s ∪ star s).centralizer.centralizer := by + rw [coe_centralizer, StarMemClass.star_coe_eq, Set.union_self, coe_centralizer] + end Centralizer end NonUnitalStarSubalgebra + +namespace NonUnitalStarAlgebra + +open NonUnitalStarSubalgebra + +variable [CommSemiring R] [StarRing R] +variable [NonUnitalSemiring A] [StarRing A] [Module R A] +variable [IsScalarTower R A A] [SMulCommClass R A A] [StarModule R A] + +variable (R) in +lemma adjoin_le_centralizer_centralizer (s : Set A) : + adjoin R s ≤ centralizer R (centralizer R s) := by + rw [← toNonUnitalSubalgebra_le_iff, centralizer_toNonUnitalSubalgebra, + adjoin_toNonUnitalSubalgebra] + convert NonUnitalAlgebra.adjoin_le_centralizer_centralizer R (s ∪ star s) + rw [StarMemClass.star_coe_eq] + simp + +lemma commute_of_mem_adjoin_of_forall_mem_commute {a b : A} {s : Set A} + (hb : b ∈ adjoin R s) (h : ∀ b ∈ s, Commute a b) (h_star : ∀ b ∈ s, Commute a (star b)) : + Commute a b := + NonUnitalAlgebra.commute_of_mem_adjoin_of_forall_mem_commute hb fun b hb ↦ + hb.elim (h b) (by simpa using h_star (star b)) + +lemma commute_of_mem_adjoin_singleton_of_commute {a b c : A} + (hc : c ∈ adjoin R {b}) (h : Commute a b) (h_star : Commute a (star b)) : + Commute a c := + commute_of_mem_adjoin_of_forall_mem_commute hc (by simpa) (by simpa) + +lemma commute_of_mem_adjoin_self {a b : A} [IsStarNormal a] (hb : b ∈ adjoin R {a}) : + Commute a b := + commute_of_mem_adjoin_singleton_of_commute hb rfl (isStarNormal_iff a |>.mp inferInstance).symm + +variable (R) in +/-- If all elements of `s : Set A` commute pairwise and with elements of `star s`, then `adjoin R s` +is a non-unital commutative semiring. + +See note [reducible non-instances]. -/ +abbrev adjoinNonUnitalCommSemiringOfComm {s : Set A} (hcomm : ∀ a ∈ s, ∀ b ∈ s, a * b = b * a) + (hcomm_star : ∀ a ∈ s, ∀ b ∈ s, a * star b = star b * a) : + NonUnitalCommSemiring (adjoin R s) := + { (adjoin R s).toNonUnitalSemiring with + mul_comm := fun ⟨_, h₁⟩ ⟨_, h₂⟩ ↦ by + have hcomm : ∀ a ∈ s ∪ star s, ∀ b ∈ s ∪ star s, a * b = b * a := fun a ha b hb ↦ + Set.union_star_self_comm (fun _ ha _ hb ↦ hcomm _ hb _ ha) + (fun _ ha _ hb ↦ hcomm_star _ hb _ ha) b hb a ha + have := adjoin_le_centralizer_centralizer R s + apply this at h₁ + apply this at h₂ + rw [← SetLike.mem_coe, coe_centralizer_centralizer] at h₁ h₂ + exact Subtype.ext <| Set.centralizer_centralizer_comm_of_comm hcomm _ h₁ _ h₂ } + +/-- If all elements of `s : Set A` commute pairwise and with elements of `star s`, then `adjoin R s` +is a non-unital commutative ring. + +See note [reducible non-instances]. -/ +abbrev adjoinNonUnitalCommRingOfComm (R : Type*) {A : Type*} [CommRing R] [StarRing R] + [NonUnitalRing A] [StarRing A] [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] + [StarModule R A] {s : Set A} (hcomm : ∀ a ∈ s, ∀ b ∈ s, a * b = b * a) + (hcomm_star : ∀ a ∈ s, ∀ b ∈ s, a * star b = star b * a) : NonUnitalCommRing (adjoin R s) := + { (adjoin R s).toNonUnitalRing, adjoinNonUnitalCommSemiringOfComm R hcomm hcomm_star with } + +end NonUnitalStarAlgebra From 0e836d6e1a3c5ed008688622e261e19fbef05e0e Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 20 Nov 2024 06:53:43 +0000 Subject: [PATCH 128/829] feat: `maintainer merge?`/`maintainer delegate?` switch for spoiler (#19230) --- .github/workflows/maintainer_merge.yml | 3 ++- scripts/maintainer_merge_message.sh | 17 ++++++++++++----- 2 files changed, 14 insertions(+), 6 deletions(-) diff --git a/.github/workflows/maintainer_merge.yml b/.github/workflows/maintainer_merge.yml index 87ab7583093b7..8be16fcfc4ed6 100644 --- a/.github/workflows/maintainer_merge.yml +++ b/.github/workflows/maintainer_merge.yml @@ -42,7 +42,8 @@ jobs: printf 'Comment:"%s"\n' "${COMMENT}" m_or_d="$(printf '%s' "${COMMENT}" | - sed -n 's=^maintainer *\(merge\|delegate\) *$=\1=p' | head -1)" + # captures `maintainer merge/delegate` as well as an optional `?`, ignoring subsequent spaces + sed -n 's=^maintainer *\(merge\|delegate\)\(?\?\) *$=\1\2=p' | head -1)" printf $'"maintainer delegate" or "maintainer merge" found? \'%s\'\n' "${m_or_d}" diff --git a/scripts/maintainer_merge_message.sh b/scripts/maintainer_merge_message.sh index ec4e14808bf86..0a8290705dd42 100755 --- a/scripts/maintainer_merge_message.sh +++ b/scripts/maintainer_merge_message.sh @@ -36,10 +36,17 @@ esac >&2 echo "EVENT_NAME: '${EVENT_NAME}'" >&2 printf 'COMMENT\n%s\nEND COMMENT\n' "${PR_COMMENT}" -# replace backticks in the title with single quotes -unbacktickedTitle="${PR_TITLE//\`/\'}" +printf '%s requested a maintainer **%s** from %s on PR [#%s](%s):\n' "${AUTHOR}" "${M_or_D/$'?'/}" "${SOURCE}" "${PR}" "${URL}" +# if `maintainer merge/delegate` is followed by `?`, then print a `spoiler` with the full comment +if [ ${M_or_D: -1} == $'?' ] +then + # replace backticks in the title with single quotes + unbacktickedTitle="${PR_TITLE//\`/\'}" ->&2 echo "neat title: '${unbacktickedTitle}'" + >&2 echo "neat title: '${unbacktickedTitle}'" -printf '%s requested a maintainer **%s** from %s on PR [#%s](%s):\n' "${AUTHOR}" "${M_or_D}" "${SOURCE}" "${PR}" "${URL}" -printf '```spoiler %s\n%s\n```\n' "${unbacktickedTitle}" "${PR_COMMENT}" + printf '```spoiler %s\n%s\n```\n' "${unbacktickedTitle}" "${PR_COMMENT}" +# otherwise, just print the title of the PR +else + printf '> %s\n' "${PR_TITLE}" +fi From 995c25cb9962967328d2ac33a0f46ebfe7a3d897 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 20 Nov 2024 08:15:52 +0000 Subject: [PATCH 129/829] feat(Data/Nat/Nth): `nth_mem_anti` (#19169) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add another lemma for working with `Nat.nth`: ```lean lemma nth_mem_anti {a b : ℕ} (hab : a ≤ b) (h : p (nth p b)) : p (nth p a) := by ``` Feel free to golf if you see a simpler way of proving this. --- Mathlib/Data/Nat/Nth.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/Mathlib/Data/Nat/Nth.lean b/Mathlib/Data/Nat/Nth.lean index fe2658ba4d808..4cde210602f99 100644 --- a/Mathlib/Data/Nat/Nth.lean +++ b/Mathlib/Data/Nat/Nth.lean @@ -252,6 +252,18 @@ theorem le_nth_of_lt_nth_succ {k a : ℕ} (h : a < nth p (k + 1)) (ha : p a) : a · rcases subset_range_nth ha with ⟨n, rfl⟩ rwa [nth_lt_nth hf, Nat.lt_succ_iff, ← nth_le_nth hf] at h +lemma nth_mem_anti {a b : ℕ} (hab : a ≤ b) (h : p (nth p b)) : p (nth p a) := by + by_cases h' : ∀ hf : (setOf p).Finite, a < #hf.toFinset + · exact nth_mem a h' + · simp only [not_forall, not_lt] at h' + have h'b : ∃ hf : (setOf p).Finite, #hf.toFinset ≤ b := by + rcases h' with ⟨hf, ha⟩ + exact ⟨hf, ha.trans hab⟩ + have ha0 : nth p a = 0 := by simp [nth_eq_zero, h'] + have hb0 : nth p b = 0 := by simp [nth_eq_zero, h'b] + rw [ha0] + rwa [hb0] at h + section Count variable (p) [DecidablePred p] From 23f454fb99cf9de12d57d1640ee2b4bc86b25413 Mon Sep 17 00:00:00 2001 From: D-Thomine <100795491+D-Thomine@users.noreply.github.com> Date: Wed, 20 Nov 2024 08:25:05 +0000 Subject: [PATCH 130/829] refactor(EReal): add add/sub lemmas and refactor limsup_add on EReal (#18879) This PR : - adds a few lemmas about addition/subtraction in `Topology.Instances.EReal`. These are translations in `EReal` of similar lemmas about multiplication/division in `ENNReal` (see `Data.ENNReal.Inv`). - renames [EReal.add_le_of_forall_add_le](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Data/Real/EReal.html#EReal.add_le_of_forall_add_le), [EReal.le_add_of_forall_le_add](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Data/Real/EReal.html#EReal.le_add_of_forall_le_add) to bring them in line with similar lemmas in `Real`/`ENNReal` (see e.g. [add_le_of_forall_lt](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Algebra/Order/Group/DenselyOrdered.html#add_le_of_forall_lt)) - simplifies significantly the proofs of `add_le_of_forall_lt` and `le_add_of_forall_lt`. - renames [EReal.add_liminf_le_liminf_add](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Topology/Instances/EReal.html#EReal.add_liminf_le_liminf_add), [EReal.limsup_add_le_add_limsup](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Topology/Instances/EReal.html#EReal.limsup_add_le_add_limsup), [EReal.limsup_add_liminf_le_limsup_add](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Topology/Instances/EReal.html#EReal.limsup_add_liminf_le_limsup_add), [EReal.liminf_add_le_limsup_add_liminf](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Topology/Instances/EReal.html#EReal.liminf_add_le_limsup_add_liminf) to bring them in line with similar lemmas in `Real`/`ENNReal` (see e.g. [ENNReal.limsup_mul_le](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/Filter/ENNReal.html#ENNReal.limsup_mul_le)) - removes [EReal.le_of_forall_lt_iff_le](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Data/Real/EReal.html#EReal.le_of_forall_lt_iff_le), [EReal.ge_of_forall_gt_iff_ge](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Data/Real/EReal.html#EReal.ge_of_forall_gt_iff_ge), [EReal.limsup_le_iff](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Topology/Instances/EReal.html#EReal.limsup_le_iff) which were very specialized lemmas used only in the older version of the proofs above. --- Mathlib/Data/Real/EReal.lean | 176 +++++++++++------- .../TopologicalEntropy/CoverEntropy.lean | 2 +- .../Integral/VitaliCaratheodory.lean | 2 +- Mathlib/Topology/Instances/EReal.lean | 57 +++--- 4 files changed, 138 insertions(+), 99 deletions(-) diff --git a/Mathlib/Data/Real/EReal.lean b/Mathlib/Data/Real/EReal.lean index 4727e2cbbcc78..723fbe073737f 100644 --- a/Mathlib/Data/Real/EReal.lean +++ b/Mathlib/Data/Real/EReal.lean @@ -923,10 +923,16 @@ def negOrderIso : EReal ≃o ERealᵒᵈ := invFun := fun x => -OrderDual.ofDual x map_rel_iff' := neg_le_neg_iff } -theorem neg_lt_iff_neg_lt {a b : EReal} : -a < b ↔ -b < a := by +theorem neg_lt_comm {a b : EReal} : -a < b ↔ -b < a := by rw [← neg_lt_neg_iff, neg_neg] + +@[deprecated (since := "2024-11-19")] alias neg_lt_iff_neg_lt := neg_lt_comm + +theorem neg_lt_of_neg_lt {a b : EReal} (h : -a < b) : -b < a := neg_lt_comm.1 h + +theorem lt_neg_comm {a b : EReal} : a < -b ↔ b < -a := by rw [← neg_lt_neg_iff, neg_neg] -theorem neg_lt_of_neg_lt {a b : EReal} (h : -a < b) : -b < a := neg_lt_iff_neg_lt.1 h +theorem lt_neg_of_lt_neg {a b : EReal} (h : a < -b) : b < -a := lt_neg_comm.1 h lemma neg_add {x y : EReal} (h1 : x ≠ ⊥ ∨ y ≠ ⊤) (h2 : x ≠ ⊤ ∨ y ≠ ⊥) : - (x + y) = - x - y := by @@ -937,64 +943,6 @@ lemma neg_sub {x y : EReal} (h1 : x ≠ ⊥ ∨ y ≠ ⊥) (h2 : x ≠ ⊤ ∨ y - (x - y) = - x + y := by rw [sub_eq_add_neg, neg_add _ _, sub_eq_add_neg, neg_neg] <;> simp_all -/-! ### Addition and order -/ - -lemma le_of_forall_lt_iff_le {x y : EReal} : (∀ z : ℝ, x < z → y ≤ z) ↔ y ≤ x := by - refine ⟨fun h ↦ WithBot.le_of_forall_lt_iff_le.1 ?_, fun h _ x_z ↦ h.trans x_z.le⟩ - rw [WithTop.forall] - aesop - -lemma ge_of_forall_gt_iff_ge {x y : EReal} : (∀ z : ℝ, z < y → z ≤ x) ↔ y ≤ x := by - refine ⟨fun h ↦ WithBot.ge_of_forall_gt_iff_ge.1 ?_, fun h _ x_z ↦ x_z.le.trans h⟩ - rw [WithTop.forall] - aesop - -/-- This lemma is superseded by `add_le_of_forall_add_le`. -/ -private lemma top_add_le_of_forall_add_le {a b : EReal} (h : ∀ c < ⊤, ∀ d < a, c + d ≤ b) : - ⊤ + a ≤ b := by - induction a with - | h_bot => exact add_bot ⊤ ▸ bot_le - | h_real a => - refine top_add_coe a ▸ le_of_forall_lt_iff_le.1 fun c b_c ↦ ?_ - specialize h (c - a + 1) (coe_lt_top (c - a + 1)) (a - 1) - rw [← coe_one, ← coe_sub, ← coe_sub, ← coe_add, ← coe_add, add_add_sub_cancel, sub_add_cancel, - EReal.coe_lt_coe_iff] at h - exact (not_le_of_lt b_c (h (sub_one_lt a))).rec - | h_top => - refine top_add_top ▸ le_of_forall_lt_iff_le.1 fun c b_c ↦ ?_ - specialize h c (coe_lt_top c) 0 zero_lt_top - rw [add_zero] at h - exact (not_le_of_lt b_c h).rec - -lemma add_le_of_forall_add_le {a b c : EReal} (h : ∀ d < a, ∀ e < b, d + e ≤ c) : a + b ≤ c := by - induction a with - | h_bot => exact bot_add b ▸ bot_le - | h_real a => induction b with - | h_bot => exact add_bot (a : EReal) ▸ bot_le - | h_real b => - refine (@ge_of_forall_gt_iff_ge c (a+b)).1 fun d d_ab ↦ ?_ - rw [← coe_add, EReal.coe_lt_coe_iff] at d_ab - rcases exists_between d_ab with ⟨e, e_d, e_ab⟩ - have key₁ : (a + d - e : ℝ) < (a : EReal) := by apply EReal.coe_lt_coe_iff.2; linarith - have key₂ : (e - a : ℝ) < (b : EReal) := by apply EReal.coe_lt_coe_iff.2; linarith - apply le_of_eq_of_le _ (h (a + d - e) key₁ (e - a) key₂) - rw [← coe_add, ← coe_sub, ← coe_sub, ← coe_add, sub_add_sub_cancel, add_sub_cancel_left] - | h_top => - rw [add_comm (a : EReal) ⊤] - exact top_add_le_of_forall_add_le fun d d_top e e_a ↦ (add_comm d e ▸ h e e_a d d_top) - | h_top => exact top_add_le_of_forall_add_le h - -lemma le_add_of_forall_le_add {a b c : EReal} (h₁ : a ≠ ⊥ ∨ b ≠ ⊤) (h₂ : a ≠ ⊤ ∨ b ≠ ⊥) - (h : ∀ d > a, ∀ e > b, c ≤ d + e) : - c ≤ a + b := by - rw [← neg_le_neg_iff, neg_add h₁ h₂] - refine add_le_of_forall_add_le fun d d_a e e_b ↦ ?_ - have h₃ : d ≠ ⊥ ∨ e ≠ ⊤ := Or.inr (ne_top_of_lt e_b) - have h₄ : d ≠ ⊤ ∨ e ≠ ⊥ := Or.inl (ne_top_of_lt d_a) - rw [← neg_neg d, neg_lt_iff_neg_lt, neg_neg a] at d_a - rw [← neg_neg e, neg_lt_iff_neg_lt, neg_neg b] at e_b - exact le_neg_of_le_neg <| neg_add h₃ h₄ ▸ h (- d) d_a (- e) e_b - /-! ### Subtraction @@ -1023,6 +971,12 @@ theorem top_sub_coe (x : ℝ) : (⊤ : EReal) - x = ⊤ := theorem coe_sub_bot (x : ℝ) : (x : EReal) - ⊥ = ⊤ := rfl +lemma sub_bot {a : EReal} (h : a ≠ ⊥) : a - ⊥ = ⊤ := by + induction a + · simp only [ne_eq, not_true_eq_false] at h + · rw [coe_sub_bot] + · rw [top_sub_bot] + theorem sub_le_sub {x y z t : EReal} (h : x ≤ y) (h' : t ≤ z) : x - z ≤ y - t := add_le_add h (neg_le_neg_iff.2 h') @@ -1047,11 +1001,105 @@ theorem toReal_sub {x y : EReal} (hx : x ≠ ⊤) (h'x : x ≠ ⊥) (hy : y ≠ lift y to ℝ using ⟨hy, h'y⟩ rfl -lemma add_sub_cancel_right {a : EReal} {b : Real} : a + b - b = a := by +lemma sub_add_cancel_left {a : EReal} {b : Real} : a - b + b = a := by induction a - · rw [bot_add b, bot_sub b] + · rw [bot_sub b, bot_add b] · norm_cast; linarith - · rw [top_add_of_ne_bot (coe_ne_bot b), top_sub_coe] + · rw [top_sub_coe b, top_add_coe b] + +lemma add_sub_cancel_right {a : EReal} {b : Real} : a + b - b = a := by + rw [sub_eq_add_neg, add_assoc, add_comm (b : EReal), ← add_assoc, ← sub_eq_add_neg] + exact sub_add_cancel_left + +lemma le_sub_iff_add_le {a b c : EReal} (hb : b ≠ ⊥ ∨ c ≠ ⊥) (ht : b ≠ ⊤ ∨ c ≠ ⊤) : + a ≤ c - b ↔ a + b ≤ c := by + induction b with + | h_bot => + simp only [ne_eq, not_true_eq_false, false_or] at hb + simp only [sub_bot hb, le_top, add_bot, bot_le] + | h_real b => + rw [← (addLECancellable_coe b).add_le_add_iff_right, sub_add_cancel_left] + | h_top => + simp only [ne_eq, not_true_eq_false, false_or, sub_top, le_bot_iff] at ht ⊢ + refine ⟨fun h ↦ h ▸ (bot_add ⊤).symm ▸ bot_le, fun h ↦ ?_⟩ + by_contra ha + exact (h.trans_lt (Ne.lt_top ht)).ne (add_top_iff_ne_bot.2 ha) + +lemma sub_le_iff_le_add {a b c : EReal} (h₁ : b ≠ ⊥ ∨ c ≠ ⊤) (h₂ : b ≠ ⊤ ∨ c ≠ ⊥) : + a - b ≤ c ↔ a ≤ c + b := by + suffices a + (-b) ≤ c ↔ a ≤ c - (-b) by simpa [sub_eq_add_neg] + refine (le_sub_iff_add_le ?_ ?_).symm <;> simpa + +protected theorem lt_sub_iff_add_lt {a b c : EReal} (h₁ : b ≠ ⊥ ∨ c ≠ ⊤) (h₂ : b ≠ ⊤ ∨ c ≠ ⊥) : + c < a - b ↔ c + b < a := + lt_iff_lt_of_le_iff_le (sub_le_iff_le_add h₁ h₂) + +theorem sub_le_of_le_add {a b c : EReal} (h : a ≤ b + c) : a - c ≤ b := by + induction c with + | h_bot => rw [add_bot, le_bot_iff] at h; simp only [h, bot_sub, bot_le] + | h_real c => exact (sub_le_iff_le_add (.inl (coe_ne_bot c)) (.inl (coe_ne_top c))).2 h + | h_top => simp only [sub_top, bot_le] + +/-- See also `EReal.sub_le_of_le_add`.-/ +theorem sub_le_of_le_add' {a b c : EReal} (h : a ≤ b + c) : a - b ≤ c := + sub_le_of_le_add (add_comm b c ▸ h) + +lemma add_le_of_le_sub {a b c : EReal} (h : a ≤ b - c) : a + c ≤ b := by + rw [← neg_neg c] + exact sub_le_of_le_add h + +lemma sub_lt_iff {a b c : EReal} (h₁ : b ≠ ⊥ ∨ c ≠ ⊥) (h₂ : b ≠ ⊤ ∨ c ≠ ⊤) : + c - b < a ↔ c < a + b := + lt_iff_lt_of_le_iff_le (le_sub_iff_add_le h₁ h₂) + +lemma add_lt_of_lt_sub {a b c : EReal} (h : a < b - c) : a + c < b := by + contrapose! h + exact sub_le_of_le_add h + +lemma sub_lt_of_lt_add {a b c : EReal} (h : a < b + c) : a - c < b := + add_lt_of_lt_sub <| by rwa [sub_eq_add_neg, neg_neg] + +/-- See also `EReal.sub_lt_of_lt_add`.-/ +lemma sub_lt_of_lt_add' {a b c : EReal} (h : a < b + c) : a - b < c := + sub_lt_of_lt_add <| by rwa [add_comm] + +/-! ### Addition and order -/ + +lemma le_of_forall_lt_iff_le {x y : EReal} : (∀ z : ℝ, x < z → y ≤ z) ↔ y ≤ x := by + refine ⟨fun h ↦ WithBot.le_of_forall_lt_iff_le.1 ?_, fun h _ x_z ↦ h.trans x_z.le⟩ + rw [WithTop.forall] + aesop + +lemma ge_of_forall_gt_iff_ge {x y : EReal} : (∀ z : ℝ, z < y → z ≤ x) ↔ y ≤ x := by + refine ⟨fun h ↦ WithBot.ge_of_forall_gt_iff_ge.1 ?_, fun h _ x_z ↦ x_z.le.trans h⟩ + rw [WithTop.forall] + aesop + +private lemma exists_lt_add_left {a b c : EReal} (hc : c < a + b) : ∃ a' < a, c < a' + b := by + obtain ⟨a', hc', ha'⟩ := exists_between (sub_lt_of_lt_add hc) + refine ⟨a', ha', (sub_lt_iff (.inl ?_) (.inr hc.ne_top)).1 hc'⟩ + contrapose! hc + exact hc ▸ (add_bot a).symm ▸ bot_le + +private lemma exists_lt_add_right {a b c : EReal} (hc : c < a + b) : ∃ b' < b, c < a + b' := by + simp_rw [add_comm a] at hc ⊢; exact exists_lt_add_left hc + +lemma add_le_of_forall_lt {a b c : EReal} (h : ∀ a' < a, ∀ b' < b, a' + b' ≤ c) : a + b ≤ c := by + refine le_of_forall_ge_of_dense fun d hd ↦ ?_ + obtain ⟨a', ha', hd⟩ := exists_lt_add_left hd + obtain ⟨b', hb', hd⟩ := exists_lt_add_right hd + exact hd.le.trans (h _ ha' _ hb') + +lemma le_add_of_forall_gt {a b c : EReal} (h₁ : a ≠ ⊥ ∨ b ≠ ⊤) (h₂ : a ≠ ⊤ ∨ b ≠ ⊥) + (h : ∀ a' > a, ∀ b' > b, c ≤ a' + b') : c ≤ a + b := by + rw [← neg_le_neg_iff, neg_add h₁ h₂] + exact add_le_of_forall_lt fun a' ha' b' hb' ↦ le_neg_of_le_neg + <| (h (-a') (lt_neg_of_lt_neg ha') (-b') (lt_neg_of_lt_neg hb')).trans_eq + (neg_add (.inr hb'.ne_top) (.inl ha'.ne_top)).symm + +@[deprecated (since := "2024-11-19")] alias top_add_le_of_forall_add_le := add_le_of_forall_lt +@[deprecated (since := "2024-11-19")] alias add_le_of_forall_add_le := add_le_of_forall_lt +@[deprecated (since := "2024-11-19")] alias le_add_of_forall_le_add := le_add_of_forall_gt lemma _root_.ENNReal.toEReal_sub {x y : ℝ≥0∞} (hy_top : y ≠ ∞) (h_le : y ≤ x) : (x - y).toEReal = x.toEReal - y.toEReal := by @@ -1576,7 +1624,7 @@ lemma div_right_distrib_of_nonneg {a b c : EReal} (h : 0 ≤ a) (h' : 0 ≤ b) : (a + b) / c = (a / c) + (b / c) := EReal.right_distrib_of_nonneg h h' -/-! #### Division and Order s -/ +/-! #### Division and Order -/ lemma monotone_div_right_of_nonneg {b : EReal} (h : 0 ≤ b) : Monotone fun a ↦ a / b := fun _ _ h' ↦ mul_le_mul_of_nonneg_right h' (inv_nonneg_of_nonneg h) diff --git a/Mathlib/Dynamics/TopologicalEntropy/CoverEntropy.lean b/Mathlib/Dynamics/TopologicalEntropy/CoverEntropy.lean index b5c4368d3eecd..f12dd6c32afc6 100644 --- a/Mathlib/Dynamics/TopologicalEntropy/CoverEntropy.lean +++ b/Mathlib/Dynamics/TopologicalEntropy/CoverEntropy.lean @@ -483,7 +483,7 @@ lemma coverEntropyEntourage_le_log_coverMincard_div {T : X → X} {F : Set X} (F eventually_atTop.2 ⟨1, fun m m_pos ↦ log_coverMincard_le_add F_inv U_symm n_pos m_pos⟩ apply ((limsup_le_limsup) key).trans suffices h : atTop.limsup v = 0 by - have := @limsup_add_le_add_limsup ℕ atTop u v + have := @limsup_add_le ℕ atTop u v rw [h, add_zero] at this specialize this (Or.inr EReal.zero_ne_top) (Or.inr EReal.zero_ne_bot) exact this.trans_eq (limsup_const (log (coverMincard T F U n) / n)) diff --git a/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean b/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean index 2aa0a2e287af9..2365cd675affa 100644 --- a/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean +++ b/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean @@ -527,7 +527,7 @@ theorem exists_upperSemicontinuous_lt_integral_gt [SigmaFinite μ] (f : α → rcases exists_lt_lowerSemicontinuous_integral_lt (fun x => -f x) hf.neg εpos with ⟨g, g_lt_f, gcont, g_integrable, g_lt_top, gint⟩ refine ⟨fun x => -g x, ?_, ?_, ?_, ?_, ?_⟩ - · exact fun x => EReal.neg_lt_iff_neg_lt.1 (by simpa only [EReal.coe_neg] using g_lt_f x) + · exact fun x => EReal.neg_lt_comm.1 (by simpa only [EReal.coe_neg] using g_lt_f x) · exact continuous_neg.comp_lowerSemicontinuous_antitone gcont fun x y hxy => EReal.neg_le_neg_iff.2 hxy diff --git a/Mathlib/Topology/Instances/EReal.lean b/Mathlib/Topology/Instances/EReal.lean index cedc98e15dfa0..746c9fd365e5b 100644 --- a/Mathlib/Topology/Instances/EReal.lean +++ b/Mathlib/Topology/Instances/EReal.lean @@ -214,65 +214,56 @@ lemma liminf_neg : liminf (- v) f = - limsup v f := lemma limsup_neg : limsup (- v) f = - liminf v f := EReal.negOrderIso.liminf_apply.symm -lemma add_liminf_le_liminf_add : (liminf u f) + (liminf v f) ≤ liminf (u + v) f := by - refine add_le_of_forall_add_le fun a a_u b b_v ↦ (le_liminf_iff).2 fun c c_ab ↦ ?_ +lemma le_liminf_add : (liminf u f) + (liminf v f) ≤ liminf (u + v) f := by + refine add_le_of_forall_lt fun a a_u b b_v ↦ (le_liminf_iff).2 fun c c_ab ↦ ?_ filter_upwards [eventually_lt_of_lt_liminf a_u, eventually_lt_of_lt_liminf b_v] with x a_x b_x - exact lt_trans c_ab (add_lt_add a_x b_x) + exact c_ab.trans (add_lt_add a_x b_x) -lemma limsup_add_le_add_limsup (h : limsup u f ≠ ⊥ ∨ limsup v f ≠ ⊤) - (h' : limsup u f ≠ ⊤ ∨ limsup v f ≠ ⊥) : +lemma limsup_add_le (h : limsup u f ≠ ⊥ ∨ limsup v f ≠ ⊤) (h' : limsup u f ≠ ⊤ ∨ limsup v f ≠ ⊥) : limsup (u + v) f ≤ (limsup u f) + (limsup v f) := by - refine le_add_of_forall_le_add h h' fun a a_u b b_v ↦ (limsup_le_iff).2 fun c c_ab ↦ ?_ + refine le_add_of_forall_gt h h' fun a a_u b b_v ↦ (limsup_le_iff).2 fun c c_ab ↦ ?_ filter_upwards [eventually_lt_of_limsup_lt a_u, eventually_lt_of_limsup_lt b_v] with x a_x b_x exact (add_lt_add a_x b_x).trans c_ab -lemma limsup_add_liminf_le_limsup_add : (limsup u f) + (liminf v f) ≤ limsup (u + v) f := - add_le_of_forall_add_le fun _ a_u _ b_v ↦ (le_limsup_iff).2 fun _ c_ab ↦ - Frequently.mono (Frequently.and_eventually ((frequently_lt_of_lt_limsup) a_u) - ((eventually_lt_of_lt_liminf) b_v)) fun _ ab_x ↦ c_ab.trans (add_lt_add ab_x.1 ab_x.2) +lemma le_limsup_add : (limsup u f) + (liminf v f) ≤ limsup (u + v) f := + add_le_of_forall_lt fun _ a_u _ b_v ↦ (le_limsup_iff).2 fun _ c_ab ↦ + (((frequently_lt_of_lt_limsup) a_u).and_eventually ((eventually_lt_of_lt_liminf) b_v)).mono + fun _ ab_x ↦ c_ab.trans (add_lt_add ab_x.1 ab_x.2) -lemma liminf_add_le_limsup_add_liminf (h : limsup u f ≠ ⊥ ∨ liminf v f ≠ ⊤) - (h' : limsup u f ≠ ⊤ ∨ liminf v f ≠ ⊥) : +lemma liminf_add_le (h : limsup u f ≠ ⊥ ∨ liminf v f ≠ ⊤) (h' : limsup u f ≠ ⊤ ∨ liminf v f ≠ ⊥) : liminf (u + v) f ≤ (limsup u f) + (liminf v f) := - le_add_of_forall_le_add h h' fun _ a_u _ b_v ↦ (liminf_le_iff).2 fun _ c_ab ↦ - Frequently.mono (Frequently.and_eventually ((frequently_lt_of_liminf_lt) b_v) - ((eventually_lt_of_limsup_lt) a_u)) fun _ ab_x ↦ (add_lt_add ab_x.2 ab_x.1).trans c_ab + le_add_of_forall_gt h h' fun _ a_u _ b_v ↦ (liminf_le_iff).2 fun _ c_ab ↦ + (((frequently_lt_of_liminf_lt) b_v).and_eventually ((eventually_lt_of_limsup_lt) a_u)).mono + fun _ ab_x ↦ (add_lt_add ab_x.2 ab_x.1).trans c_ab + +@[deprecated (since := "2024-11-11")] alias add_liminf_le_liminf_add := le_liminf_add +@[deprecated (since := "2024-11-11")] alias limsup_add_le_add_limsup := limsup_add_le +@[deprecated (since := "2024-11-11")] alias limsup_add_liminf_le_limsup_add := le_limsup_add +@[deprecated (since := "2024-11-11")] alias liminf_add_le_limsup_add_liminf := liminf_add_le variable {a b : EReal} lemma limsup_add_bot_of_ne_top (h : limsup u f = ⊥) (h' : limsup v f ≠ ⊤) : limsup (u + v) f = ⊥ := by - apply le_bot_iff.1 (le_trans (limsup_add_le_add_limsup (Or.inr h') _) _) - · rw [h]; exact Or.inl bot_ne_top + apply le_bot_iff.1 ((limsup_add_le (.inr h') _).trans _) + · rw [h]; exact .inl bot_ne_top · rw [h, bot_add] lemma limsup_add_le_of_le (ha : limsup u f < a) (hb : limsup v f ≤ b) : limsup (u + v) f ≤ a + b := by - rcases eq_top_or_lt_top b with (rfl | h) + rcases eq_top_or_lt_top b with rfl | h · rw [add_top_of_ne_bot ha.ne_bot]; exact le_top - · exact le_trans (limsup_add_le_add_limsup (Or.inr (lt_of_le_of_lt hb h).ne) (Or.inl ha.ne_top)) - (add_le_add ha.le hb) + · exact (limsup_add_le (.inr (hb.trans_lt h).ne) (.inl ha.ne_top)).trans (add_le_add ha.le hb) lemma liminf_add_gt_of_gt (ha : a < liminf u f) (hb : b < liminf v f) : a + b < liminf (u + v) f := - lt_of_lt_of_le (add_lt_add ha hb) add_liminf_le_liminf_add + (add_lt_add ha hb).trans_le le_liminf_add lemma liminf_add_top_of_ne_bot (h : liminf u f = ⊤) (h' : liminf v f ≠ ⊥) : liminf (u + v) f = ⊤ := by - apply top_le_iff.1 (le_trans _ (add_liminf_le_liminf_add)) + apply top_le_iff.1 (le_trans _ le_liminf_add) rw [h, top_add_of_ne_bot h'] -lemma limsup_le_iff {b : EReal} : limsup u f ≤ b ↔ ∀ c : ℝ, b < c → ∀ᶠ a : α in f, u a ≤ c := by - rw [← le_of_forall_lt_iff_le] - refine ⟨?_, ?_⟩ <;> intro h c b_c - · rcases exists_between_coe_real b_c with ⟨d, b_d, d_c⟩ - apply mem_of_superset (eventually_lt_of_limsup_lt (lt_of_le_of_lt (h d b_d) d_c)) - rw [Set.setOf_subset_setOf] - exact fun _ h' ↦ h'.le - · rcases eq_or_neBot f with rfl | _ - · simp only [limsup_bot, bot_le] - · exact (limsup_le_of_le) (h c b_c) - end LimInfSup /-! ### Continuity of addition -/ From 55c020e1b67b5ae4dfe3efcc1f71a0346ef31a9e Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Wed, 20 Nov 2024 08:25:07 +0000 Subject: [PATCH 131/829] feat(Order/WithBot): Equiv.withBotSubtypeNe (#19251) Needed in #19210 From the Carleson project --- Mathlib/Order/WithBot.lean | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/Mathlib/Order/WithBot.lean b/Mathlib/Order/WithBot.lean index d49c2f822a4fb..5be2bf90e1157 100644 --- a/Mathlib/Order/WithBot.lean +++ b/Mathlib/Order/WithBot.lean @@ -158,6 +158,13 @@ theorem eq_unbot_iff {a : α} {b : WithBot α} (h : b ≠ ⊥) : · simpa using h rfl · simp +/-- The equivalence between the non-bottom elements of `WithBot α` and `α`. -/ +@[simps] def _root_.Equiv.withBotSubtypeNe : {y : WithBot α // y ≠ ⊥} ≃ α where + toFun := fun ⟨x,h⟩ => WithBot.unbot x h + invFun x := ⟨x, WithBot.coe_ne_bot⟩ + left_inv _ := by simp + right_inv _ := by simp + section LE variable [LE α] @@ -729,6 +736,13 @@ theorem eq_untop_iff {a : α} {b : WithTop α} (h : b ≠ ⊤) : a = b.untop h ↔ a = b := WithBot.eq_unbot_iff (α := αᵒᵈ) h +/-- The equivalence between the non-top elements of `WithTop α` and `α`. -/ +@[simps] def _root_.Equiv.withTopSubtypeNe : {y : WithTop α // y ≠ ⊤} ≃ α where + toFun := fun ⟨x,h⟩ => WithTop.untop x h + invFun x := ⟨x, WithTop.coe_ne_top⟩ + left_inv _ := by simp + right_inv _:= by simp + section LE variable [LE α] From 650a3f818f1cd3ac8de8856c884065a6d1f6749e Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Wed, 20 Nov 2024 09:05:39 +0000 Subject: [PATCH 132/829] chore: no `open Foo` after `namespace Foo` (#19223) ...and not opening the same name space twice. Co-authored-by: Moritz Firsching --- Mathlib/Algebra/CubicDiscriminant.lean | 2 -- Mathlib/Algebra/MvPolynomial/Cardinal.lean | 2 -- Mathlib/Algebra/Polynomial/Identities.lean | 2 -- Mathlib/Algebra/Polynomial/Mirror.lean | 2 -- Mathlib/Algebra/Polynomial/Monomial.lean | 2 -- Mathlib/Algebra/Polynomial/Reverse.lean | 4 ++-- Mathlib/Algebra/Polynomial/Splits.lean | 2 -- Mathlib/Algebra/Polynomial/Taylor.lean | 2 -- Mathlib/Analysis/Normed/Operator/BanachSteinhaus.lean | 2 -- Mathlib/Analysis/RCLike/Lemmas.lean | 2 -- .../Analysis/SpecialFunctions/Trigonometric/Basic.lean | 2 -- Mathlib/CategoryTheory/Closed/Cartesian.lean | 2 +- Mathlib/CategoryTheory/FiberedCategory/BasedCategory.lean | 2 +- Mathlib/CategoryTheory/Sites/Grothendieck.lean | 2 +- Mathlib/CategoryTheory/Sites/PreservesSheafification.lean | 2 +- Mathlib/CategoryTheory/Sites/Pretopology.lean | 2 +- Mathlib/Combinatorics/Enumerative/Catalan.lean | 2 -- Mathlib/Data/Matrix/Basis.lean | 2 -- Mathlib/Data/Matrix/Hadamard.lean | 2 -- Mathlib/Data/Matrix/Kronecker.lean | 2 -- Mathlib/Data/Nat/Choose/Dvd.lean | 2 -- Mathlib/Data/Nat/MaxPowDiv.lean | 2 -- Mathlib/Data/Nat/Periodic.lean | 2 +- Mathlib/Data/Rat/Lemmas.lean | 2 -- Mathlib/Data/Real/Irrational.lean | 2 -- Mathlib/FieldTheory/IntermediateField/Basic.lean | 2 -- Mathlib/LinearAlgebra/CrossProduct.lean | 2 -- Mathlib/LinearAlgebra/Matrix/Charpoly/Minpoly.lean | 2 -- Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean | 2 -- Mathlib/LinearAlgebra/Matrix/Symmetric.lean | 2 -- Mathlib/LinearAlgebra/Matrix/Transvection.lean | 2 -- Mathlib/LinearAlgebra/QuadraticForm/Prod.lean | 2 -- .../MeasureTheory/Function/StronglyMeasurable/Basic.lean | 2 -- .../NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean | 2 -- Mathlib/NumberTheory/Primorial.lean | 2 -- Mathlib/RingTheory/Adjoin/PowerBasis.lean | 2 -- Mathlib/RingTheory/EisensteinCriterion.lean | 2 -- Mathlib/RingTheory/IntegralDomain.lean | 2 -- Mathlib/RingTheory/Jacobson.lean | 2 -- Mathlib/RingTheory/MatrixAlgebra.lean | 8 +------- Mathlib/RingTheory/Nullstellensatz.lean | 2 -- Mathlib/RingTheory/Polynomial/Content.lean | 2 -- Mathlib/RingTheory/Polynomial/Dickson.lean | 2 -- Mathlib/RingTheory/Polynomial/Pochhammer.lean | 2 -- Mathlib/RingTheory/Polynomial/ScaleRoots.lean | 2 -- Mathlib/RingTheory/ReesAlgebra.lean | 2 -- Mathlib/SetTheory/Cardinal/Divisibility.lean | 2 -- Mathlib/Topology/Algebra/Nonarchimedean/Basic.lean | 2 -- Mathlib/Topology/Algebra/Polynomial.lean | 2 -- MathlibTest/matrix.lean | 2 -- MathlibTest/norm_cast.lean | 2 -- 51 files changed, 9 insertions(+), 101 deletions(-) diff --git a/Mathlib/Algebra/CubicDiscriminant.lean b/Mathlib/Algebra/CubicDiscriminant.lean index 42e56fed93eff..4104d2add1285 100644 --- a/Mathlib/Algebra/CubicDiscriminant.lean +++ b/Mathlib/Algebra/CubicDiscriminant.lean @@ -40,8 +40,6 @@ structure Cubic (R : Type*) where namespace Cubic -open Cubic Polynomial - open Polynomial variable {R S F K : Type*} diff --git a/Mathlib/Algebra/MvPolynomial/Cardinal.lean b/Mathlib/Algebra/MvPolynomial/Cardinal.lean index 880c15e0cedd3..3ee977de57b85 100644 --- a/Mathlib/Algebra/MvPolynomial/Cardinal.lean +++ b/Mathlib/Algebra/MvPolynomial/Cardinal.lean @@ -20,8 +20,6 @@ universe u v open Cardinal -open Cardinal - namespace MvPolynomial section TwoUniverses diff --git a/Mathlib/Algebra/Polynomial/Identities.lean b/Mathlib/Algebra/Polynomial/Identities.lean index d33dfaf9c2923..3aa9df0fe8d2a 100644 --- a/Mathlib/Algebra/Polynomial/Identities.lean +++ b/Mathlib/Algebra/Polynomial/Identities.lean @@ -18,8 +18,6 @@ noncomputable section namespace Polynomial -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} diff --git a/Mathlib/Algebra/Polynomial/Mirror.lean b/Mathlib/Algebra/Polynomial/Mirror.lean index 9f3b754cd0a0a..804488801b7dc 100644 --- a/Mathlib/Algebra/Polynomial/Mirror.lean +++ b/Mathlib/Algebra/Polynomial/Mirror.lean @@ -27,8 +27,6 @@ divisible by `X`. namespace Polynomial -open Polynomial - section Semiring variable {R : Type*} [Semiring R] (p q : R[X]) diff --git a/Mathlib/Algebra/Polynomial/Monomial.lean b/Mathlib/Algebra/Polynomial/Monomial.lean index 535103330d899..f994bf2057858 100644 --- a/Mathlib/Algebra/Polynomial/Monomial.lean +++ b/Mathlib/Algebra/Polynomial/Monomial.lean @@ -16,8 +16,6 @@ noncomputable section namespace Polynomial -open Polynomial - universe u variable {R : Type u} {a b : R} {m n : ℕ} diff --git a/Mathlib/Algebra/Polynomial/Reverse.lean b/Mathlib/Algebra/Polynomial/Reverse.lean index 07d080133f5ee..a5ba40706e958 100644 --- a/Mathlib/Algebra/Polynomial/Reverse.lean +++ b/Mathlib/Algebra/Polynomial/Reverse.lean @@ -19,9 +19,9 @@ coefficients of `f` and `g` do not multiply to zero. namespace Polynomial -open Polynomial Finsupp Finset +open Finsupp Finset -open Polynomial +open scoped Polynomial section Semiring diff --git a/Mathlib/Algebra/Polynomial/Splits.lean b/Mathlib/Algebra/Polynomial/Splits.lean index 918c15651b4a4..211e8294de371 100644 --- a/Mathlib/Algebra/Polynomial/Splits.lean +++ b/Mathlib/Algebra/Polynomial/Splits.lean @@ -33,8 +33,6 @@ variable {R : Type*} {F : Type u} {K : Type v} {L : Type w} namespace Polynomial -open Polynomial - section Splits section CommRing diff --git a/Mathlib/Algebra/Polynomial/Taylor.lean b/Mathlib/Algebra/Polynomial/Taylor.lean index 805deafc3c1c7..f3f2c2b9b9fd7 100644 --- a/Mathlib/Algebra/Polynomial/Taylor.lean +++ b/Mathlib/Algebra/Polynomial/Taylor.lean @@ -26,8 +26,6 @@ noncomputable section namespace Polynomial -open Polynomial - variable {R : Type*} [Semiring R] (r : R) (f : R[X]) /-- The Taylor expansion of a polynomial `f` at `r`. -/ diff --git a/Mathlib/Analysis/Normed/Operator/BanachSteinhaus.lean b/Mathlib/Analysis/Normed/Operator/BanachSteinhaus.lean index deb7cba16f499..c112b860d9378 100644 --- a/Mathlib/Analysis/Normed/Operator/BanachSteinhaus.lean +++ b/Mathlib/Analysis/Normed/Operator/BanachSteinhaus.lean @@ -37,8 +37,6 @@ theorem banach_steinhaus {ι : Type*} [CompleteSpace E] {g : ι → E →SL[σ open ENNReal -open ENNReal - /-- This version of Banach-Steinhaus is stated in terms of suprema of `↑‖·‖₊ : ℝ≥0∞` for convenience. -/ theorem banach_steinhaus_iSup_nnnorm {ι : Type*} [CompleteSpace E] {g : ι → E →SL[σ₁₂] F} diff --git a/Mathlib/Analysis/RCLike/Lemmas.lean b/Mathlib/Analysis/RCLike/Lemmas.lean index 0603a4fda4c16..9ce1f0f40f0a5 100644 --- a/Mathlib/Analysis/RCLike/Lemmas.lean +++ b/Mathlib/Analysis/RCLike/Lemmas.lean @@ -13,8 +13,6 @@ variable {K E : Type*} [RCLike K] namespace Polynomial -open Polynomial - theorem ofReal_eval (p : ℝ[X]) (x : ℝ) : (↑(p.eval x) : K) = aeval (↑x) p := (@aeval_algebraMap_apply_eq_algebraMap_eval ℝ K _ _ _ x p).symm diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean index 68908a896942d..9dd56c0fcd556 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean @@ -201,8 +201,6 @@ end NNReal namespace Real -open Real - @[simp] theorem sin_pi : sin π = 0 := by rw [← mul_div_cancel_left₀ π (two_ne_zero' ℝ), two_mul, add_div, sin_add, cos_pi_div_two]; simp diff --git a/Mathlib/CategoryTheory/Closed/Cartesian.lean b/Mathlib/CategoryTheory/Closed/Cartesian.lean index 6b72829dd536d..0607788f02110 100644 --- a/Mathlib/CategoryTheory/Closed/Cartesian.lean +++ b/Mathlib/CategoryTheory/Closed/Cartesian.lean @@ -44,7 +44,7 @@ noncomputable section namespace CategoryTheory -open CategoryTheory Category Limits MonoidalCategory +open Category Limits MonoidalCategory /-- An object `X` is *exponentiable* if `(X × -)` is a left adjoint. We define this as being `Closed` in the cartesian monoidal structure. diff --git a/Mathlib/CategoryTheory/FiberedCategory/BasedCategory.lean b/Mathlib/CategoryTheory/FiberedCategory/BasedCategory.lean index 5319308872660..1f67ad02341c3 100644 --- a/Mathlib/CategoryTheory/FiberedCategory/BasedCategory.lean +++ b/Mathlib/CategoryTheory/FiberedCategory/BasedCategory.lean @@ -29,7 +29,7 @@ universe v₅ u₅ v₄ u₄ v₃ u₃ v₂ u₂ v₁ u₁ namespace CategoryTheory -open CategoryTheory Functor Category NatTrans IsHomLift +open Functor Category NatTrans IsHomLift variable {𝒮 : Type u₁} [Category.{v₁} 𝒮] diff --git a/Mathlib/CategoryTheory/Sites/Grothendieck.lean b/Mathlib/CategoryTheory/Sites/Grothendieck.lean index 28c48a20268de..5570569b02e2b 100644 --- a/Mathlib/CategoryTheory/Sites/Grothendieck.lean +++ b/Mathlib/CategoryTheory/Sites/Grothendieck.lean @@ -53,7 +53,7 @@ universe v₁ u₁ v u namespace CategoryTheory -open CategoryTheory Category +open Category variable (C : Type u) [Category.{v} C] diff --git a/Mathlib/CategoryTheory/Sites/PreservesSheafification.lean b/Mathlib/CategoryTheory/Sites/PreservesSheafification.lean index 52f8413c10939..2c7a0002cec28 100644 --- a/Mathlib/CategoryTheory/Sites/PreservesSheafification.lean +++ b/Mathlib/CategoryTheory/Sites/PreservesSheafification.lean @@ -43,7 +43,7 @@ universe v u namespace CategoryTheory -open CategoryTheory Category Limits +open Category Limits variable {C : Type u} [Category.{v} C] (J : GrothendieckTopology C) {A B : Type*} [Category A] [Category B] (F : A ⥤ B) diff --git a/Mathlib/CategoryTheory/Sites/Pretopology.lean b/Mathlib/CategoryTheory/Sites/Pretopology.lean index f2f7dfc06495c..d888a2611acf2 100644 --- a/Mathlib/CategoryTheory/Sites/Pretopology.lean +++ b/Mathlib/CategoryTheory/Sites/Pretopology.lean @@ -35,7 +35,7 @@ noncomputable section namespace CategoryTheory -open CategoryTheory Category Limits Presieve +open Category Limits Presieve variable {C : Type u} [Category.{v} C] [HasPullbacks C] variable (C) diff --git a/Mathlib/Combinatorics/Enumerative/Catalan.lean b/Mathlib/Combinatorics/Enumerative/Catalan.lean index c2ae512bb889f..21a578fab89bf 100644 --- a/Mathlib/Combinatorics/Enumerative/Catalan.lean +++ b/Mathlib/Combinatorics/Enumerative/Catalan.lean @@ -140,8 +140,6 @@ theorem catalan_three : catalan 3 = 5 := by namespace Tree -open Tree - /-- Given two finsets, find all trees that can be formed with left child in `a` and right child in `b` -/ abbrev pairwiseNode (a b : Finset (Tree Unit)) : Finset (Tree Unit) := diff --git a/Mathlib/Data/Matrix/Basis.lean b/Mathlib/Data/Matrix/Basis.lean index f9d9157fe8645..7c49644844bcd 100644 --- a/Mathlib/Data/Matrix/Basis.lean +++ b/Mathlib/Data/Matrix/Basis.lean @@ -19,8 +19,6 @@ variable {R α : Type*} namespace Matrix -open Matrix - variable [DecidableEq l] [DecidableEq m] [DecidableEq n] section Zero diff --git a/Mathlib/Data/Matrix/Hadamard.lean b/Mathlib/Data/Matrix/Hadamard.lean index d94a90b94b736..cdaa2360b95c5 100644 --- a/Mathlib/Data/Matrix/Hadamard.lean +++ b/Mathlib/Data/Matrix/Hadamard.lean @@ -34,8 +34,6 @@ variable {α m n R : Type*} namespace Matrix -open Matrix - /-- `Matrix.hadamard` defines the Hadamard product, which is the pointwise product of two matrices of the same size. -/ def hadamard [Mul α] (A : Matrix m n α) (B : Matrix m n α) : Matrix m n α := diff --git a/Mathlib/Data/Matrix/Kronecker.lean b/Mathlib/Data/Matrix/Kronecker.lean index bbbe63c6841d9..1d108492e89cf 100644 --- a/Mathlib/Data/Matrix/Kronecker.lean +++ b/Mathlib/Data/Matrix/Kronecker.lean @@ -43,8 +43,6 @@ These require `open Kronecker`: namespace Matrix - -open Matrix open scoped RightActions variable {R α α' β β' γ γ' : Type*} diff --git a/Mathlib/Data/Nat/Choose/Dvd.lean b/Mathlib/Data/Nat/Choose/Dvd.lean index 0272d236dd30a..37ff1fb69906d 100644 --- a/Mathlib/Data/Nat/Choose/Dvd.lean +++ b/Mathlib/Data/Nat/Choose/Dvd.lean @@ -13,8 +13,6 @@ import Mathlib.Data.Nat.Prime.Factorial namespace Nat -open Nat - namespace Prime variable {p a b k : ℕ} diff --git a/Mathlib/Data/Nat/MaxPowDiv.lean b/Mathlib/Data/Nat/MaxPowDiv.lean index 34905e1420912..c9ea1fcff18cf 100644 --- a/Mathlib/Data/Nat/MaxPowDiv.lean +++ b/Mathlib/Data/Nat/MaxPowDiv.lean @@ -21,8 +21,6 @@ The implementation of `maxPowDiv` improves on the speed of `padicValNat`. namespace Nat -open Nat - /-- Tail recursive function which returns the largest `k : ℕ` such that `p ^ k ∣ n` for any `p : ℕ`. `padicValNat_eq_maxPowDiv` allows the code generator to use this definition for `padicValNat` diff --git a/Mathlib/Data/Nat/Periodic.lean b/Mathlib/Data/Nat/Periodic.lean index 1e2e335bec976..0165a52cf2c76 100644 --- a/Mathlib/Data/Nat/Periodic.lean +++ b/Mathlib/Data/Nat/Periodic.lean @@ -18,7 +18,7 @@ periodic predicates which helps determine their cardinality when filtering inter namespace Nat -open Nat Function +open Function theorem periodic_gcd (a : ℕ) : Periodic (gcd a) a := by simp only [forall_const, gcd_add_self_right, eq_self_iff_true, Periodic] diff --git a/Mathlib/Data/Rat/Lemmas.lean b/Mathlib/Data/Rat/Lemmas.lean index f12050bbcfeaf..dd45de30bdee3 100644 --- a/Mathlib/Data/Rat/Lemmas.lean +++ b/Mathlib/Data/Rat/Lemmas.lean @@ -16,8 +16,6 @@ import Mathlib.Data.PNat.Defs namespace Rat -open Rat - theorem num_dvd (a) {b : ℤ} (b0 : b ≠ 0) : (a /. b).num ∣ a := by cases' e : a /. b with n d h c rw [Rat.mk'_eq_divInt, divInt_eq_iff b0 (mod_cast h)] at e diff --git a/Mathlib/Data/Real/Irrational.lean b/Mathlib/Data/Real/Irrational.lean index 62ea491e08723..59bde31ad0787 100644 --- a/Mathlib/Data/Real/Irrational.lean +++ b/Mathlib/Data/Real/Irrational.lean @@ -452,8 +452,6 @@ section Polynomial open Polynomial -open Polynomial - variable (x : ℝ) (p : ℤ[X]) theorem one_lt_natDegree_of_irrational_root (hx : Irrational x) (p_nonzero : p ≠ 0) diff --git a/Mathlib/FieldTheory/IntermediateField/Basic.lean b/Mathlib/FieldTheory/IntermediateField/Basic.lean index 994c57c28bf6d..33d4aa8b3bf97 100644 --- a/Mathlib/FieldTheory/IntermediateField/Basic.lean +++ b/Mathlib/FieldTheory/IntermediateField/Basic.lean @@ -37,8 +37,6 @@ intermediate field, field extension -/ -open Polynomial - open Polynomial variable (K L L' : Type*) [Field K] [Field L] [Field L'] [Algebra K L] [Algebra K L'] diff --git a/Mathlib/LinearAlgebra/CrossProduct.lean b/Mathlib/LinearAlgebra/CrossProduct.lean index 8fb6567b4ec21..3093c7a58ad50 100644 --- a/Mathlib/LinearAlgebra/CrossProduct.lean +++ b/Mathlib/LinearAlgebra/CrossProduct.lean @@ -36,8 +36,6 @@ crossproduct -/ -open Matrix - open Matrix variable {R : Type*} [CommRing R] diff --git a/Mathlib/LinearAlgebra/Matrix/Charpoly/Minpoly.lean b/Mathlib/LinearAlgebra/Matrix/Charpoly/Minpoly.lean index 7f0eb6a40db51..e2a73551493a8 100644 --- a/Mathlib/LinearAlgebra/Matrix/Charpoly/Minpoly.lean +++ b/Mathlib/LinearAlgebra/Matrix/Charpoly/Minpoly.lean @@ -28,8 +28,6 @@ open Finset namespace Matrix -open Matrix - variable (M : Matrix n n R) @[simp] diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean index dafa3eb44607d..1001a688e6203 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean @@ -42,8 +42,6 @@ open Equiv Equiv.Perm Finset Function namespace Matrix -open Matrix - variable {m n : Type*} [DecidableEq n] [Fintype n] [DecidableEq m] [Fintype m] variable {R : Type v} [CommRing R] diff --git a/Mathlib/LinearAlgebra/Matrix/Symmetric.lean b/Mathlib/LinearAlgebra/Matrix/Symmetric.lean index 83d9a2f3dccba..498e22fb3bacc 100644 --- a/Mathlib/LinearAlgebra/Matrix/Symmetric.lean +++ b/Mathlib/LinearAlgebra/Matrix/Symmetric.lean @@ -25,8 +25,6 @@ variable {α β n m R : Type*} namespace Matrix -open Matrix - /-- A matrix `A : Matrix n n α` is "symmetric" if `Aᵀ = A`. -/ def IsSymm (A : Matrix n n α) : Prop := Aᵀ = A diff --git a/Mathlib/LinearAlgebra/Matrix/Transvection.lean b/Mathlib/LinearAlgebra/Matrix/Transvection.lean index 1580376e82cc2..1c777ee3fb514 100644 --- a/Mathlib/LinearAlgebra/Matrix/Transvection.lean +++ b/Mathlib/LinearAlgebra/Matrix/Transvection.lean @@ -63,8 +63,6 @@ universe u₁ u₂ namespace Matrix -open Matrix - variable (n p : Type*) (R : Type u₂) {𝕜 : Type*} [Field 𝕜] variable [DecidableEq n] [DecidableEq p] variable [CommRing R] diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean b/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean index 328b2319e0765..c9ef8e47cf38d 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean @@ -37,8 +37,6 @@ variable {ι : Type*} {R : Type*} {M₁ M₂ N₁ N₂ P : Type*} {Mᵢ Nᵢ : namespace QuadraticMap -open QuadraticMap - section Prod section Semiring diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean index 7ab91eabbf17b..b2da38021c0f9 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean @@ -512,8 +512,6 @@ variable [MeasurableSpace α] [TopologicalSpace β] open Filter -open Filter - @[aesop safe 20 (rule_sets := [Measurable])] protected theorem sup [Max β] [ContinuousSup β] (hf : StronglyMeasurable f) (hg : StronglyMeasurable g) : StronglyMeasurable (f ⊔ g) := diff --git a/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean b/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean index fabb26655f44b..6d1c1c2f19d9b 100644 --- a/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean +++ b/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean @@ -23,8 +23,6 @@ use to show the class number of the ring of integers of a function field is fini namespace Polynomial -open Polynomial - open AbsoluteValue Real variable {Fq : Type*} [Fintype Fq] diff --git a/Mathlib/NumberTheory/Primorial.lean b/Mathlib/NumberTheory/Primorial.lean index 3867a3aa0f893..d2a7922592c84 100644 --- a/Mathlib/NumberTheory/Primorial.lean +++ b/Mathlib/NumberTheory/Primorial.lean @@ -27,8 +27,6 @@ open Finset open Nat -open Nat - /-- The primorial `n#` of `n` is the product of the primes less than or equal to `n`. -/ def primorial (n : ℕ) : ℕ := ∏ p ∈ range (n + 1) with p.Prime, p diff --git a/Mathlib/RingTheory/Adjoin/PowerBasis.lean b/Mathlib/RingTheory/Adjoin/PowerBasis.lean index 115e9eae03f82..2a3bbcac0b345 100644 --- a/Mathlib/RingTheory/Adjoin/PowerBasis.lean +++ b/Mathlib/RingTheory/Adjoin/PowerBasis.lean @@ -79,8 +79,6 @@ namespace PowerBasis open Polynomial -open Polynomial - variable {R : Type*} [CommRing R] [Algebra R S] [Algebra R K] [IsScalarTower R K S] variable {A : Type*} [CommRing A] [Algebra R A] [Algebra S A] variable [IsScalarTower R S A] {B : PowerBasis S A} diff --git a/Mathlib/RingTheory/EisensteinCriterion.lean b/Mathlib/RingTheory/EisensteinCriterion.lean index 216ad0c0c0a4d..0e246b9233e41 100644 --- a/Mathlib/RingTheory/EisensteinCriterion.lean +++ b/Mathlib/RingTheory/EisensteinCriterion.lean @@ -22,8 +22,6 @@ variable {R : Type*} [CommRing R] namespace Polynomial -open Polynomial - namespace EisensteinCriterionAux -- Section for auxiliary lemmas used in the proof of `irreducible_of_eisenstein_criterion` diff --git a/Mathlib/RingTheory/IntegralDomain.lean b/Mathlib/RingTheory/IntegralDomain.lean index b292e8b1c1ffc..a44906bcb806f 100644 --- a/Mathlib/RingTheory/IntegralDomain.lean +++ b/Mathlib/RingTheory/IntegralDomain.lean @@ -149,8 +149,6 @@ section EuclideanDivision namespace Polynomial -open Polynomial - variable (K : Type) [Field K] [Algebra R[X] K] [IsFractionRing R[X] K] theorem div_eq_quo_add_rem_div (f : R[X]) {g : R[X]} (hg : g.Monic) : diff --git a/Mathlib/RingTheory/Jacobson.lean b/Mathlib/RingTheory/Jacobson.lean index 14e36f678fd84..56997e889c66f 100644 --- a/Mathlib/RingTheory/Jacobson.lean +++ b/Mathlib/RingTheory/Jacobson.lean @@ -247,8 +247,6 @@ end Localization namespace Polynomial -open Polynomial - section CommRing -- Porting note: move to better place diff --git a/Mathlib/RingTheory/MatrixAlgebra.lean b/Mathlib/RingTheory/MatrixAlgebra.lean index 0750d23424344..8fcf5dc5f622c 100644 --- a/Mathlib/RingTheory/MatrixAlgebra.lean +++ b/Mathlib/RingTheory/MatrixAlgebra.lean @@ -14,13 +14,7 @@ suppress_compilation universe u v w -open TensorProduct - -open TensorProduct - -open Algebra.TensorProduct - -open Matrix +open TensorProduct Algebra.TensorProduct Matrix variable {R : Type u} [CommSemiring R] variable {A : Type v} [Semiring A] [Algebra R A] diff --git a/Mathlib/RingTheory/Nullstellensatz.lean b/Mathlib/RingTheory/Nullstellensatz.lean index f28499af60736..71d381cfeb340 100644 --- a/Mathlib/RingTheory/Nullstellensatz.lean +++ b/Mathlib/RingTheory/Nullstellensatz.lean @@ -28,8 +28,6 @@ noncomputable section namespace MvPolynomial -open MvPolynomial - variable {k : Type*} [Field k] variable {σ : Type*} diff --git a/Mathlib/RingTheory/Polynomial/Content.lean b/Mathlib/RingTheory/Polynomial/Content.lean index c92ff552d810e..394d392597e7b 100644 --- a/Mathlib/RingTheory/Polynomial/Content.lean +++ b/Mathlib/RingTheory/Polynomial/Content.lean @@ -30,8 +30,6 @@ Let `p : R[X]`. namespace Polynomial -open Polynomial - section Primitive variable {R : Type*} [CommSemiring R] diff --git a/Mathlib/RingTheory/Polynomial/Dickson.lean b/Mathlib/RingTheory/Polynomial/Dickson.lean index d5d51a3dee789..1872e6ac8a9e2 100644 --- a/Mathlib/RingTheory/Polynomial/Dickson.lean +++ b/Mathlib/RingTheory/Polynomial/Dickson.lean @@ -52,8 +52,6 @@ noncomputable section namespace Polynomial -open Polynomial - variable {R S : Type*} [CommRing R] [CommRing S] (k : ℕ) (a : R) /-- `dickson` is the `n`-th (generalised) Dickson polynomial of the `k`-th kind associated to the diff --git a/Mathlib/RingTheory/Polynomial/Pochhammer.lean b/Mathlib/RingTheory/Polynomial/Pochhammer.lean index 63708f272d1ed..d6872afb27cbc 100644 --- a/Mathlib/RingTheory/Polynomial/Pochhammer.lean +++ b/Mathlib/RingTheory/Polynomial/Pochhammer.lean @@ -38,8 +38,6 @@ universe u v open Polynomial -open Polynomial - section Semiring variable (S : Type u) [Semiring S] diff --git a/Mathlib/RingTheory/Polynomial/ScaleRoots.lean b/Mathlib/RingTheory/Polynomial/ScaleRoots.lean index dce2e27c49300..38ee5b2e41344 100644 --- a/Mathlib/RingTheory/Polynomial/ScaleRoots.lean +++ b/Mathlib/RingTheory/Polynomial/ScaleRoots.lean @@ -20,8 +20,6 @@ variable {R S A K : Type*} namespace Polynomial -open Polynomial - section Semiring variable [Semiring R] [Semiring S] diff --git a/Mathlib/RingTheory/ReesAlgebra.lean b/Mathlib/RingTheory/ReesAlgebra.lean index 127232d147d6d..2069a0781786a 100644 --- a/Mathlib/RingTheory/ReesAlgebra.lean +++ b/Mathlib/RingTheory/ReesAlgebra.lean @@ -30,8 +30,6 @@ variable {R M : Type u} [CommRing R] [AddCommGroup M] [Module R M] (I : Ideal R) open Polynomial -open Polynomial - /-- The Rees algebra of an ideal `I`, defined as the subalgebra of `R[X]` whose `i`-th coefficient falls in `I ^ i`. -/ def reesAlgebra : Subalgebra R R[X] where diff --git a/Mathlib/SetTheory/Cardinal/Divisibility.lean b/Mathlib/SetTheory/Cardinal/Divisibility.lean index 6a176f3246b65..a82232c834d80 100644 --- a/Mathlib/SetTheory/Cardinal/Divisibility.lean +++ b/Mathlib/SetTheory/Cardinal/Divisibility.lean @@ -31,8 +31,6 @@ Note furthermore that no infinite cardinal is irreducible namespace Cardinal -open Cardinal - universe u variable {a b : Cardinal.{u}} {n m : ℕ} diff --git a/Mathlib/Topology/Algebra/Nonarchimedean/Basic.lean b/Mathlib/Topology/Algebra/Nonarchimedean/Basic.lean index 4f3a24fdd08bc..d339eba913c78 100644 --- a/Mathlib/Topology/Algebra/Nonarchimedean/Basic.lean +++ b/Mathlib/Topology/Algebra/Nonarchimedean/Basic.lean @@ -105,8 +105,6 @@ end NonarchimedeanGroup namespace NonarchimedeanRing -open NonarchimedeanRing - open NonarchimedeanAddGroup variable {R S : Type*} diff --git a/Mathlib/Topology/Algebra/Polynomial.lean b/Mathlib/Topology/Algebra/Polynomial.lean index 663b8bd81dc20..cda8bf62a44a1 100644 --- a/Mathlib/Topology/Algebra/Polynomial.lean +++ b/Mathlib/Topology/Algebra/Polynomial.lean @@ -37,8 +37,6 @@ open IsAbsoluteValue Filter namespace Polynomial -open Polynomial - section TopologicalSemiring variable {R S : Type*} [Semiring R] [TopologicalSpace R] [TopologicalSemiring R] (p : R[X]) diff --git a/MathlibTest/matrix.lean b/MathlibTest/matrix.lean index d8571679b3a9c..f6e84240f8d07 100644 --- a/MathlibTest/matrix.lean +++ b/MathlibTest/matrix.lean @@ -13,8 +13,6 @@ variable {α β : Type} [Semiring α] [Ring β] namespace Matrix -open Matrix - /-! Test that the dimensions are inferred correctly, even for empty matrices -/ section dimensions diff --git a/MathlibTest/norm_cast.lean b/MathlibTest/norm_cast.lean index cb792593c1ef5..2cd0529030d98 100644 --- a/MathlibTest/norm_cast.lean +++ b/MathlibTest/norm_cast.lean @@ -126,8 +126,6 @@ example (a b : ℕ) (h2 : ((a + b + 0 : ℕ) : ℤ) = 10) : -- example {x : ℚ} : ((x + 42 : ℚ) : ℝ) = x + 42 := by push_cast namespace ENNReal - -open ENNReal lemma half_lt_self_bis {a : ℝ≥0∞} (hz : a ≠ 0) (ht : a ≠ ⊤) : a / 2 < a := by lift a to NNReal using ht have h : (2 : ℝ≥0∞) = ((2 : NNReal) : ℝ≥0∞) := rfl From ef950f1e9fabe40ab2cc6447154355a295c85a74 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Wed, 20 Nov 2024 09:05:41 +0000 Subject: [PATCH 133/829] chore: delete unused private lemma in Complex/exponential (#19263) --- Mathlib/Data/Complex/Exponential.lean | 5 ----- 1 file changed, 5 deletions(-) diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index 87be1b9c712d2..1054c7a046c4b 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -456,9 +456,6 @@ theorem cos_zero : cos 0 = 1 := by simp [cos] @[simp] theorem cos_neg : cos (-x) = cos x := by simp [cos, sub_eq_add_neg, exp_neg, add_comm] -private theorem cos_add_aux {a b c d : ℂ} : - (a + b) * (c + d) - (b - a) * (d - c) * -1 = 2 * (a * c + b * d) := by ring - theorem cos_add : cos (x + y) = cos x * cos y - sin x * sin y := by rw [← cosh_mul_I, add_mul, cosh_add, cosh_mul_I, cosh_mul_I, sinh_mul_I, sinh_mul_I, mul_mul_mul_comm, I_mul_I, mul_neg_one, sub_eq_add_neg] @@ -1498,5 +1495,3 @@ theorem abs_exp_eq_iff_re_eq {x y : ℂ} : abs (exp x) = abs (exp y) ↔ x.re = rw [abs_exp, abs_exp, Real.exp_eq_exp] end Complex - -set_option linter.style.longFile 1700 From 2fb27be908b6ad6a43a55a68946651eb967332f1 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Wed, 20 Nov 2024 09:38:40 +0000 Subject: [PATCH 134/829] chore(Algebra/Module): further split `Defs.lean` (#18995) The goal is that we can define `Module` without depending on `Units`. The following moves out of `Defs.lean` are included in this PR: * Module structure over Nat and Int go to `Module/NatInt.lean` * Composition with `RingHom`s goes to `Module/RingHom.lean` * Results about units go to `Module/Basic.lean` --- Mathlib.lean | 2 + Mathlib/Algebra/AddConstMap/Basic.lean | 2 +- Mathlib/Algebra/CharZero/Quotient.lean | 2 +- .../Algebra/Group/Submonoid/Pointwise.lean | 1 + Mathlib/Algebra/Homology/TotalComplex.lean | 1 + Mathlib/Algebra/Module/Basic.lean | 8 +- Mathlib/Algebra/Module/Defs.lean | 164 +----------------- Mathlib/Algebra/Module/End.lean | 1 + Mathlib/Algebra/Module/Hom.lean | 1 + Mathlib/Algebra/Module/LinearMap/Defs.lean | 2 + Mathlib/Algebra/Module/NatInt.lean | 146 ++++++++++++++++ Mathlib/Algebra/Module/RingHom.lean | 92 ++++++++++ Mathlib/Algebra/Order/Nonneg/Module.lean | 2 +- Mathlib/Algebra/Ring/Subsemiring/Basic.lean | 2 +- Mathlib/CategoryTheory/Preadditive/Basic.lean | 1 + .../SimpleGraph/Regularity/Energy.lean | 2 +- Mathlib/Data/Finsupp/Pointwise.lean | 1 + Mathlib/Data/Int/AbsoluteValue.lean | 2 +- Mathlib/Data/Matrix/Diagonal.lean | 2 + Mathlib/Data/Set/Pointwise/SMul.lean | 1 + Mathlib/GroupTheory/ArchimedeanDensely.lean | 5 +- Mathlib/GroupTheory/Divisible.lean | 1 + Mathlib/GroupTheory/FreeAbelianGroup.lean | 1 + Mathlib/GroupTheory/OrderOfElement.lean | 1 + Mathlib/GroupTheory/QuotientGroup/Basic.lean | 1 + Mathlib/Order/Filter/Germ/Basic.lean | 3 +- 26 files changed, 281 insertions(+), 166 deletions(-) create mode 100644 Mathlib/Algebra/Module/NatInt.lean create mode 100644 Mathlib/Algebra/Module/RingHom.lean diff --git a/Mathlib.lean b/Mathlib.lean index 3ca25acb72351..797b5848ce70d 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -520,6 +520,7 @@ import Mathlib.Algebra.Module.LocalizedModule.Int import Mathlib.Algebra.Module.LocalizedModule.IsLocalization import Mathlib.Algebra.Module.LocalizedModule.Submodule import Mathlib.Algebra.Module.MinimalAxioms +import Mathlib.Algebra.Module.NatInt import Mathlib.Algebra.Module.Opposite import Mathlib.Algebra.Module.PID import Mathlib.Algebra.Module.Pi @@ -533,6 +534,7 @@ import Mathlib.Algebra.Module.Presentation.Tautological import Mathlib.Algebra.Module.Prod import Mathlib.Algebra.Module.Projective import Mathlib.Algebra.Module.Rat +import Mathlib.Algebra.Module.RingHom import Mathlib.Algebra.Module.SnakeLemma import Mathlib.Algebra.Module.Submodule.Basic import Mathlib.Algebra.Module.Submodule.Bilinear diff --git a/Mathlib/Algebra/AddConstMap/Basic.lean b/Mathlib/Algebra/AddConstMap/Basic.lean index 0bf23cde46517..7dc43d9c08e9a 100644 --- a/Mathlib/Algebra/AddConstMap/Basic.lean +++ b/Mathlib/Algebra/AddConstMap/Basic.lean @@ -5,7 +5,7 @@ Authors: Yury Kudryashov -/ import Mathlib.Algebra.Group.Action.Pi import Mathlib.Algebra.GroupPower.IterateHom -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.Module.NatInt import Mathlib.Algebra.Order.Archimedean.Basic import Mathlib.Algebra.Order.Group.Instances import Mathlib.Logic.Function.Iterate diff --git a/Mathlib/Algebra/CharZero/Quotient.lean b/Mathlib/Algebra/CharZero/Quotient.lean index a0e14ca270865..d3872fea29c7e 100644 --- a/Mathlib/Algebra/CharZero/Quotient.lean +++ b/Mathlib/Algebra/CharZero/Quotient.lean @@ -5,7 +5,7 @@ Authors: Eric Wieser -/ import Mathlib.Algebra.Field.Basic import Mathlib.Algebra.Order.Group.Unbundled.Int -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.Module.NatInt import Mathlib.GroupTheory.QuotientGroup.Defs import Mathlib.Algebra.Group.Subgroup.ZPowers.Basic diff --git a/Mathlib/Algebra/Group/Submonoid/Pointwise.lean b/Mathlib/Algebra/Group/Submonoid/Pointwise.lean index 2c5e53bea3b80..9c5e50a98eb6b 100644 --- a/Mathlib/Algebra/Group/Submonoid/Pointwise.lean +++ b/Mathlib/Algebra/Group/Submonoid/Pointwise.lean @@ -5,6 +5,7 @@ Authors: Eric Wieser -/ import Mathlib.Algebra.Group.Hom.End import Mathlib.Algebra.Group.Submonoid.Membership +import Mathlib.Algebra.GroupWithZero.Action.End import Mathlib.Algebra.Order.BigOperators.Group.List import Mathlib.Data.Set.Pointwise.SMul import Mathlib.Order.WellFoundedSet diff --git a/Mathlib/Algebra/Homology/TotalComplex.lean b/Mathlib/Algebra/Homology/TotalComplex.lean index 645d747de064f..15137d2628a9b 100644 --- a/Mathlib/Algebra/Homology/TotalComplex.lean +++ b/Mathlib/Algebra/Homology/TotalComplex.lean @@ -6,6 +6,7 @@ Authors: Joël Riou import Mathlib.CategoryTheory.Linear.Basic import Mathlib.Algebra.Homology.ComplexShapeSigns import Mathlib.Algebra.Homology.HomologicalBicomplex +import Mathlib.Algebra.Module.Basic /-! # The total complex of a bicomplex diff --git a/Mathlib/Algebra/Module/Basic.lean b/Mathlib/Algebra/Module/Basic.lean index 0eeee9934e3d3..b653c02aa9f6c 100644 --- a/Mathlib/Algebra/Module/Basic.lean +++ b/Mathlib/Algebra/Module/Basic.lean @@ -6,7 +6,8 @@ Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro import Mathlib.Algebra.Field.Basic import Mathlib.Algebra.Group.Action.Pi import Mathlib.Algebra.Group.Indicator -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.GroupWithZero.Action.Units +import Mathlib.Algebra.Module.NatInt import Mathlib.Algebra.NoZeroSMulDivisors.Defs /-! @@ -23,6 +24,11 @@ universe u v variable {α R M M₂ : Type*} +@[simp] +theorem Units.neg_smul [Ring R] [AddCommGroup M] [Module R M] (u : Rˣ) (x : M) : + -u • x = -(u • x) := by + rw [Units.smul_def, Units.val_neg, _root_.neg_smul, Units.smul_def] + @[simp] theorem invOf_two_smul_add_invOf_two_smul (R) [Semiring R] [AddCommMonoid M] [Module R M] [Invertible (2 : R)] (x : M) : diff --git a/Mathlib/Algebra/Module/Defs.lean b/Mathlib/Algebra/Module/Defs.lean index c614c036f0977..ebc9ab4254580 100644 --- a/Mathlib/Algebra/Module/Defs.lean +++ b/Mathlib/Algebra/Module/Defs.lean @@ -3,10 +3,11 @@ Copyright (c) 2015 Nathaniel Thomas. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro -/ -import Mathlib.Algebra.GroupWithZero.Action.End -import Mathlib.Algebra.GroupWithZero.Action.Units +import Mathlib.Algebra.Group.Action.End +import Mathlib.Algebra.Group.Equiv.Basic +import Mathlib.Algebra.GroupWithZero.Action.Defs +import Mathlib.Algebra.Ring.Defs import Mathlib.Algebra.SMulWithZero -import Mathlib.Data.Int.Cast.Lemmas /-! # Modules over a ring @@ -39,7 +40,9 @@ assert_not_exists Field assert_not_exists Invertible assert_not_exists Multiset assert_not_exists Pi.single_smul₀ +assert_not_exists RingHom assert_not_exists Set.indicator +assert_not_exists Units open Function Set @@ -72,14 +75,6 @@ instance (priority := 100) Module.toMulActionWithZero smul_zero := smul_zero zero_smul := Module.zero_smul } -instance AddCommGroup.toNatModule : Module ℕ M where - one_smul := one_nsmul - mul_smul m n a := mul_nsmul' a m n - smul_add n a b := nsmul_add a b n - smul_zero := nsmul_zero - zero_smul := zero_nsmul - add_smul r s x := add_nsmul x r s - theorem add_smul : (r + s) • x = r • x + s • x := Module.add_smul r s x @@ -111,31 +106,7 @@ protected abbrev Function.Surjective.module [AddCommMonoid M₂] [SMul R M₂] ( rcases hf x with ⟨x, rfl⟩ rw [← f.map_zero, ← smul, zero_smul] } -/-- Push forward the action of `R` on `M` along a compatible surjective map `f : R →+* S`. - -See also `Function.Surjective.mulActionLeft` and `Function.Surjective.distribMulActionLeft`. --/ -abbrev Function.Surjective.moduleLeft {R S M : Type*} [Semiring R] [AddCommMonoid M] [Module R M] - [Semiring S] [SMul S M] (f : R →+* S) (hf : Function.Surjective f) - (hsmul : ∀ (c) (x : M), f c • x = c • x) : Module S M := - { hf.distribMulActionLeft f.toMonoidHom hsmul with - zero_smul := fun x => by rw [← f.map_zero, hsmul, zero_smul] - add_smul := hf.forall₂.mpr fun a b x => by simp only [← f.map_add, hsmul, add_smul] } - -variable {R} (M) - -/-- Compose a `Module` with a `RingHom`, with action `f s • m`. - -See note [reducible non-instances]. -/ -abbrev Module.compHom [Semiring S] (f : S →+* R) : Module S M := - { MulActionWithZero.compHom M f.toMonoidWithZeroHom, DistribMulAction.compHom M (f : S →* R) with - -- Porting note: the `show f (r + s) • x = f r • x + f s • x` wasn't needed in mathlib3. - -- Somehow, now that `SMul` is heterogeneous, it can't unfold earlier fields of a definition for - -- use in later fields. See - -- https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Heterogeneous.20scalar.20multiplication - add_smul := fun r s x => show f (r + s) • x = f r • x + f s • x by simp [add_smul] } - -variable {M} +variable {R} theorem Module.eq_zero_of_zero_eq_one (zero_eq_one : (0 : R) = 1) : x = 0 := by rw [← one_smul R x, ← zero_eq_one, zero_smul] @@ -146,20 +117,9 @@ theorem smul_add_one_sub_smul {R : Type*} [Ring R] [Module R M] {r : R} {m : M} end AddCommMonoid - section AddCommGroup -variable (R M) [Semiring R] [AddCommGroup M] - -instance AddCommGroup.toIntModule : Module ℤ M where - one_smul := one_zsmul - mul_smul m n a := mul_zsmul a m n - smul_add n a b := zsmul_add a b n - smul_zero := zsmul_zero - zero_smul := zero_zsmul - add_smul r s x := add_zsmul x r s - -variable {R M} +variable [Semiring R] [AddCommGroup M] theorem Convex.combo_eq_smul_sub_add [Module R M] {x y : M} {a b : R} (h : a + b = 1) : a • x + b • y = b • (y - x) + x := @@ -187,10 +147,6 @@ theorem neg_smul : -r • x = -(r • x) := theorem neg_smul_neg : -r • -x = r • x := by rw [neg_smul, smul_neg, neg_neg] -@[simp] -theorem Units.neg_smul (u : Rˣ) (x : M) : -u • x = -(u • x) := by - rw [Units.smul_def, Units.val_neg, _root_.neg_smul, Units.smul_def] - variable (R) theorem neg_one_smul (x : M) : (-1 : R) • x = -x := by simp @@ -202,26 +158,6 @@ theorem sub_smul (r s : R) (y : M) : (r - s) • y = r • y - s • y := by end Module -variable (R) - -/-- An `AddCommMonoid` that is a `Module` over a `Ring` carries a natural `AddCommGroup` -structure. -See note [reducible non-instances]. -/ -abbrev Module.addCommMonoidToAddCommGroup - [Ring R] [AddCommMonoid M] [Module R M] : AddCommGroup M := - { (inferInstance : AddCommMonoid M) with - neg := fun a => (-1 : R) • a - neg_add_cancel := fun a => - show (-1 : R) • a + a = 0 by - nth_rw 2 [← one_smul R a] - rw [← add_smul, neg_add_cancel, zero_smul] - zsmul := fun z a => (z : R) • a - zsmul_zero' := fun a => by simpa only [Int.cast_zero] using zero_smul R a - zsmul_succ' := fun z a => by simp [add_comm, add_smul] - zsmul_neg' := fun z a => by simp [← smul_assoc, neg_one_smul] } - -variable {R} - /-- A module over a `Subsingleton` semiring is a `Subsingleton`. We cannot register this as an instance because Lean has no way to guess `R`. -/ protected theorem Module.subsingleton (R M : Type*) [Semiring R] [Subsingleton R] [AddCommMonoid M] @@ -242,87 +178,3 @@ instance (priority := 910) Semiring.toModule [Semiring R] : Module R R where instance [NonUnitalNonAssocSemiring R] : DistribSMul R R where smul_add := left_distrib - -/-- A ring homomorphism `f : R →+* M` defines a module structure by `r • x = f r * x`. -/ -def RingHom.toModule [Semiring R] [Semiring S] (f : R →+* S) : Module R S := - Module.compHom S f - -/-- If the module action of `R` on `S` is compatible with multiplication on `S`, then -`fun x ↦ x • 1` is a ring homomorphism from `R` to `S`. - -This is the `RingHom` version of `MonoidHom.smulOneHom`. - -When `R` is commutative, usually `algebraMap` should be preferred. -/ -@[simps!] def RingHom.smulOneHom - [Semiring R] [NonAssocSemiring S] [Module R S] [IsScalarTower R S S] : R →+* S where - __ := MonoidHom.smulOneHom - map_zero' := zero_smul R 1 - map_add' := (add_smul · · 1) - -/-- A homomorphism between semirings R and S can be equivalently specified by a R-module -structure on S such that S/S/R is a scalar tower. -/ -def ringHomEquivModuleIsScalarTower [Semiring R] [Semiring S] : - (R →+* S) ≃ {_inst : Module R S // IsScalarTower R S S} where - toFun f := ⟨Module.compHom S f, SMul.comp.isScalarTower _⟩ - invFun := fun ⟨_, _⟩ ↦ RingHom.smulOneHom - left_inv f := RingHom.ext fun r ↦ mul_one (f r) - right_inv := fun ⟨_, _⟩ ↦ Subtype.ext <| Module.ext <| funext₂ <| smul_one_smul S - -section AddCommMonoid - -variable [Semiring R] [AddCommMonoid M] [Module R M] - -section - -variable (R) - -/-- `nsmul` is equal to any other module structure via a cast. -/ -lemma Nat.cast_smul_eq_nsmul (n : ℕ) (b : M) : (n : R) • b = n • b := by - induction n with - | zero => rw [Nat.cast_zero, zero_smul, zero_smul] - | succ n ih => rw [Nat.cast_succ, add_smul, add_smul, one_smul, ih, one_smul] - -/-- `nsmul` is equal to any other module structure via a cast. -/ --- See note [no_index around OfNat.ofNat] -lemma ofNat_smul_eq_nsmul (n : ℕ) [n.AtLeastTwo] (b : M) : - (no_index (OfNat.ofNat n) : R) • b = OfNat.ofNat n • b := Nat.cast_smul_eq_nsmul .. - -/-- `nsmul` is equal to any other module structure via a cast. -/ -@[deprecated Nat.cast_smul_eq_nsmul (since := "2024-07-23")] -lemma nsmul_eq_smul_cast (n : ℕ) (b : M) : n • b = (n : R) • b := (Nat.cast_smul_eq_nsmul ..).symm - -end - -/-- Convert back any exotic `ℕ`-smul to the canonical instance. This should not be needed since in -mathlib all `AddCommMonoid`s should normally have exactly one `ℕ`-module structure by design. --/ -theorem nat_smul_eq_nsmul (h : Module ℕ M) (n : ℕ) (x : M) : @SMul.smul ℕ M h.toSMul n x = n • x := - Nat.cast_smul_eq_nsmul .. - -/-- All `ℕ`-module structures are equal. Not an instance since in mathlib all `AddCommMonoid` -should normally have exactly one `ℕ`-module structure by design. -/ -def AddCommMonoid.uniqueNatModule : Unique (Module ℕ M) where - default := by infer_instance - uniq P := (Module.ext' P _) fun n => by convert nat_smul_eq_nsmul P n - -instance AddCommMonoid.nat_isScalarTower : IsScalarTower ℕ R M where - smul_assoc n x y := by - induction n with - | zero => simp only [zero_smul] - | succ n ih => simp only [add_smul, one_smul, ih] - -end AddCommMonoid - -theorem map_natCast_smul [AddCommMonoid M] [AddCommMonoid M₂] {F : Type*} [FunLike F M M₂] - [AddMonoidHomClass F M M₂] (f : F) (R S : Type*) [Semiring R] [Semiring S] [Module R M] - [Module S M₂] (x : ℕ) (a : M) : f ((x : R) • a) = (x : S) • f a := by - simp only [Nat.cast_smul_eq_nsmul, AddMonoidHom.map_nsmul, map_nsmul] - -theorem Nat.smul_one_eq_cast {R : Type*} [NonAssocSemiring R] (m : ℕ) : m • (1 : R) = ↑m := by - rw [nsmul_eq_mul, mul_one] - -theorem Int.smul_one_eq_cast {R : Type*} [NonAssocRing R] (m : ℤ) : m • (1 : R) = ↑m := by - rw [zsmul_eq_mul, mul_one] - -@[deprecated (since := "2024-05-03")] alias Nat.smul_one_eq_coe := Nat.smul_one_eq_cast -@[deprecated (since := "2024-05-03")] alias Int.smul_one_eq_coe := Int.smul_one_eq_cast diff --git a/Mathlib/Algebra/Module/End.lean b/Mathlib/Algebra/Module/End.lean index b1269c071ad3d..405cb931c87f9 100644 --- a/Mathlib/Algebra/Module/End.lean +++ b/Mathlib/Algebra/Module/End.lean @@ -5,6 +5,7 @@ Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro -/ import Mathlib.Algebra.Group.Hom.End import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.Module.NatInt /-! # Module structure and endomorphisms diff --git a/Mathlib/Algebra/Module/Hom.lean b/Mathlib/Algebra/Module/Hom.lean index 66bdfade35658..8149ac54e6b58 100644 --- a/Mathlib/Algebra/Module/Hom.lean +++ b/Mathlib/Algebra/Module/Hom.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ import Mathlib.Algebra.Group.Hom.Instances +import Mathlib.Algebra.GroupWithZero.Action.End import Mathlib.Algebra.Module.End import Mathlib.Algebra.Ring.Opposite import Mathlib.GroupTheory.GroupAction.DomAct.Basic diff --git a/Mathlib/Algebra/Module/LinearMap/Defs.lean b/Mathlib/Algebra/Module/LinearMap/Defs.lean index 893c0c84ca370..c6436036fb7c1 100644 --- a/Mathlib/Algebra/Module/LinearMap/Defs.lean +++ b/Mathlib/Algebra/Module/LinearMap/Defs.lean @@ -5,6 +5,8 @@ Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro, Anne Frédéric Dupuis, Heather Macbeth -/ import Mathlib.Algebra.Group.Hom.Instances +import Mathlib.Algebra.Module.NatInt +import Mathlib.Algebra.Module.RingHom import Mathlib.Algebra.Ring.CompTypeclasses import Mathlib.GroupTheory.GroupAction.Hom diff --git a/Mathlib/Algebra/Module/NatInt.lean b/Mathlib/Algebra/Module/NatInt.lean new file mode 100644 index 0000000000000..d1aab097a652d --- /dev/null +++ b/Mathlib/Algebra/Module/NatInt.lean @@ -0,0 +1,146 @@ +/- +Copyright (c) 2015 Nathaniel Thomas. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro +-/ +import Mathlib.Algebra.Module.Defs +import Mathlib.Data.Int.Cast.Lemmas + +/-! +# Modules over `ℕ` and `ℤ` + +This file concerns modules where the scalars are the natural numbers or the integers. + +## Main definitions + +* `AddCommGroup.toNatModule`: any `AddCommMonoid` is (uniquely) a module over the naturals. + TODO: this name is not right! +* `AddCommGroup.toIntModule`: any `AddCommGroup` is a module over the integers. + +## Main results + + * `AddCommMonoid.uniqueNatModule`: there is an unique `AddCommMonoid ℕ M` structure for any `M` + +## Tags + +semimodule, module, vector space +-/ + +assert_not_exists Field +assert_not_exists Invertible +assert_not_exists Multiset +assert_not_exists Pi.single_smul₀ +assert_not_exists Set.indicator + +open Function Set + +universe u v + +variable {R S M M₂ : Type*} + +section AddCommMonoid + +variable [Semiring R] [AddCommMonoid M] [Module R M] (r s : R) (x : M) + +instance AddCommGroup.toNatModule : Module ℕ M where + one_smul := one_nsmul + mul_smul m n a := mul_nsmul' a m n + smul_add n a b := nsmul_add a b n + smul_zero := nsmul_zero + zero_smul := zero_nsmul + add_smul r s x := add_nsmul x r s + +end AddCommMonoid + +section AddCommGroup + +variable (R M) [Semiring R] [AddCommGroup M] + +instance AddCommGroup.toIntModule : Module ℤ M where + one_smul := one_zsmul + mul_smul m n a := mul_zsmul a m n + smul_add n a b := zsmul_add a b n + smul_zero := zsmul_zero + zero_smul := zero_zsmul + add_smul r s x := add_zsmul x r s + +end AddCommGroup + +variable (R) + +/-- An `AddCommMonoid` that is a `Module` over a `Ring` carries a natural `AddCommGroup` +structure. +See note [reducible non-instances]. -/ +abbrev Module.addCommMonoidToAddCommGroup + [Ring R] [AddCommMonoid M] [Module R M] : AddCommGroup M := + { (inferInstance : AddCommMonoid M) with + neg := fun a => (-1 : R) • a + neg_add_cancel := fun a => + show (-1 : R) • a + a = 0 by + nth_rw 2 [← one_smul R a] + rw [← add_smul, neg_add_cancel, zero_smul] + zsmul := fun z a => (z : R) • a + zsmul_zero' := fun a => by simpa only [Int.cast_zero] using zero_smul R a + zsmul_succ' := fun z a => by simp [add_comm, add_smul] + zsmul_neg' := fun z a => by simp [← smul_assoc, neg_one_smul] } + +variable {R} + +section AddCommMonoid + +variable [Semiring R] [AddCommMonoid M] [Module R M] + +section + +variable (R) + +/-- `nsmul` is equal to any other module structure via a cast. -/ +lemma Nat.cast_smul_eq_nsmul (n : ℕ) (b : M) : (n : R) • b = n • b := by + induction n with + | zero => rw [Nat.cast_zero, zero_smul, zero_smul] + | succ n ih => rw [Nat.cast_succ, add_smul, add_smul, one_smul, ih, one_smul] + +/-- `nsmul` is equal to any other module structure via a cast. -/ +-- See note [no_index around OfNat.ofNat] +lemma ofNat_smul_eq_nsmul (n : ℕ) [n.AtLeastTwo] (b : M) : + (no_index (OfNat.ofNat n) : R) • b = OfNat.ofNat n • b := Nat.cast_smul_eq_nsmul .. + +/-- `nsmul` is equal to any other module structure via a cast. -/ +@[deprecated Nat.cast_smul_eq_nsmul (since := "2024-07-23")] +lemma nsmul_eq_smul_cast (n : ℕ) (b : M) : n • b = (n : R) • b := (Nat.cast_smul_eq_nsmul ..).symm + +end + +/-- Convert back any exotic `ℕ`-smul to the canonical instance. This should not be needed since in +mathlib all `AddCommMonoid`s should normally have exactly one `ℕ`-module structure by design. +-/ +theorem nat_smul_eq_nsmul (h : Module ℕ M) (n : ℕ) (x : M) : @SMul.smul ℕ M h.toSMul n x = n • x := + Nat.cast_smul_eq_nsmul .. + +/-- All `ℕ`-module structures are equal. Not an instance since in mathlib all `AddCommMonoid` +should normally have exactly one `ℕ`-module structure by design. -/ +def AddCommMonoid.uniqueNatModule : Unique (Module ℕ M) where + default := by infer_instance + uniq P := (Module.ext' P _) fun n => by convert nat_smul_eq_nsmul P n + +instance AddCommMonoid.nat_isScalarTower : IsScalarTower ℕ R M where + smul_assoc n x y := by + induction n with + | zero => simp only [zero_smul] + | succ n ih => simp only [add_smul, one_smul, ih] + +end AddCommMonoid + +theorem map_natCast_smul [AddCommMonoid M] [AddCommMonoid M₂] {F : Type*} [FunLike F M M₂] + [AddMonoidHomClass F M M₂] (f : F) (R S : Type*) [Semiring R] [Semiring S] [Module R M] + [Module S M₂] (x : ℕ) (a : M) : f ((x : R) • a) = (x : S) • f a := by + simp only [Nat.cast_smul_eq_nsmul, AddMonoidHom.map_nsmul, map_nsmul] + +theorem Nat.smul_one_eq_cast {R : Type*} [NonAssocSemiring R] (m : ℕ) : m • (1 : R) = ↑m := by + rw [nsmul_eq_mul, mul_one] + +theorem Int.smul_one_eq_cast {R : Type*} [NonAssocRing R] (m : ℤ) : m • (1 : R) = ↑m := by + rw [zsmul_eq_mul, mul_one] + +@[deprecated (since := "2024-05-03")] alias Nat.smul_one_eq_coe := Nat.smul_one_eq_cast +@[deprecated (since := "2024-05-03")] alias Int.smul_one_eq_coe := Int.smul_one_eq_cast diff --git a/Mathlib/Algebra/Module/RingHom.lean b/Mathlib/Algebra/Module/RingHom.lean new file mode 100644 index 0000000000000..ae178799ace69 --- /dev/null +++ b/Mathlib/Algebra/Module/RingHom.lean @@ -0,0 +1,92 @@ +/- +Copyright (c) 2015 Nathaniel Thomas. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro +-/ +import Mathlib.Algebra.GroupWithZero.Action.End +import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.Ring.Hom.Defs + +/-! +# Composing modules with a ring hom + +## Main definitions + + * `Module.compHom`: compose a `Module` with a `RingHom`, with action `f s • m`. + * `RingHom.toModule`: a `RingHom` defines a module structure by `r • x = f r * x`. + +## Tags + +semimodule, module, vector space +-/ + +assert_not_exists Field +assert_not_exists Invertible +assert_not_exists Multiset +assert_not_exists Pi.single_smul₀ +assert_not_exists Set.indicator + +open Function Set + +universe u v + +variable {R S M M₂ : Type*} + +section AddCommMonoid + +variable [Semiring R] [AddCommMonoid M] [Module R M] (r s : R) (x : M) + +variable (R) + +/-- Push forward the action of `R` on `M` along a compatible surjective map `f : R →+* S`. + +See also `Function.Surjective.mulActionLeft` and `Function.Surjective.distribMulActionLeft`. +-/ +abbrev Function.Surjective.moduleLeft {R S M : Type*} [Semiring R] [AddCommMonoid M] [Module R M] + [Semiring S] [SMul S M] (f : R →+* S) (hf : Function.Surjective f) + (hsmul : ∀ (c) (x : M), f c • x = c • x) : Module S M := + { hf.distribMulActionLeft f.toMonoidHom hsmul with + zero_smul := fun x => by rw [← f.map_zero, hsmul, zero_smul] + add_smul := hf.forall₂.mpr fun a b x => by simp only [← f.map_add, hsmul, add_smul] } + +variable {R} (M) + +/-- Compose a `Module` with a `RingHom`, with action `f s • m`. + +See note [reducible non-instances]. -/ +abbrev Module.compHom [Semiring S] (f : S →+* R) : Module S M := + { MulActionWithZero.compHom M f.toMonoidWithZeroHom, DistribMulAction.compHom M (f : S →* R) with + -- Porting note: the `show f (r + s) • x = f r • x + f s • x` wasn't needed in mathlib3. + -- Somehow, now that `SMul` is heterogeneous, it can't unfold earlier fields of a definition for + -- use in later fields. See + -- https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Heterogeneous.20scalar.20multiplication + add_smul := fun r s x => show f (r + s) • x = f r • x + f s • x by simp [add_smul] } + +variable {M} + +end AddCommMonoid + +/-- A ring homomorphism `f : R →+* M` defines a module structure by `r • x = f r * x`. -/ +def RingHom.toModule [Semiring R] [Semiring S] (f : R →+* S) : Module R S := + Module.compHom S f + +/-- If the module action of `R` on `S` is compatible with multiplication on `S`, then +`fun x ↦ x • 1` is a ring homomorphism from `R` to `S`. + +This is the `RingHom` version of `MonoidHom.smulOneHom`. + +When `R` is commutative, usually `algebraMap` should be preferred. -/ +@[simps!] def RingHom.smulOneHom + [Semiring R] [NonAssocSemiring S] [Module R S] [IsScalarTower R S S] : R →+* S where + __ := MonoidHom.smulOneHom + map_zero' := zero_smul R 1 + map_add' := (add_smul · · 1) + +/-- A homomorphism between semirings R and S can be equivalently specified by a R-module +structure on S such that S/S/R is a scalar tower. -/ +def ringHomEquivModuleIsScalarTower [Semiring R] [Semiring S] : + (R →+* S) ≃ {_inst : Module R S // IsScalarTower R S S} where + toFun f := ⟨Module.compHom S f, SMul.comp.isScalarTower _⟩ + invFun := fun ⟨_, _⟩ ↦ RingHom.smulOneHom + left_inv f := RingHom.ext fun r ↦ mul_one (f r) + right_inv := fun ⟨_, _⟩ ↦ Subtype.ext <| Module.ext <| funext₂ <| smul_one_smul S diff --git a/Mathlib/Algebra/Order/Nonneg/Module.lean b/Mathlib/Algebra/Order/Nonneg/Module.lean index 1a388a67825b3..fb4548137a724 100644 --- a/Mathlib/Algebra/Order/Nonneg/Module.lean +++ b/Mathlib/Algebra/Order/Nonneg/Module.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Apurva Nakade. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Apurva Nakade -/ -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.Module.RingHom import Mathlib.Algebra.Order.Module.OrderedSMul import Mathlib.Algebra.Order.Nonneg.Ring diff --git a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean index d955416892c36..bcd9bb5829c17 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Algebra.Group.Submonoid.Membership -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.Module.RingHom import Mathlib.Algebra.Ring.Action.Subobjects import Mathlib.Algebra.Ring.Equiv import Mathlib.Algebra.Ring.Prod diff --git a/Mathlib/CategoryTheory/Preadditive/Basic.lean b/Mathlib/CategoryTheory/Preadditive/Basic.lean index 0ddcb0081738b..7fade9b2835eb 100644 --- a/Mathlib/CategoryTheory/Preadditive/Basic.lean +++ b/Mathlib/CategoryTheory/Preadditive/Basic.lean @@ -5,6 +5,7 @@ Authors: Markus Himmel, Jakob von Raumer -/ import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Algebra.Group.Hom.Defs +import Mathlib.Algebra.GroupWithZero.Action.Units import Mathlib.Algebra.Module.End import Mathlib.CategoryTheory.Endomorphism import Mathlib.CategoryTheory.Limits.Shapes.Kernels diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean index 452746aebd488..ec9fd519eeeeb 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yaël Dillies, Bhavik Mehta. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Bhavik Mehta -/ -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.Module.NatInt import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Combinatorics.SimpleGraph.Density import Mathlib.Data.Rat.BigOperators diff --git a/Mathlib/Data/Finsupp/Pointwise.lean b/Mathlib/Data/Finsupp/Pointwise.lean index 059bf76ad5eed..04b6b2a36e9d5 100644 --- a/Mathlib/Data/Finsupp/Pointwise.lean +++ b/Mathlib/Data/Finsupp/Pointwise.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.Ring.InjSurj import Mathlib.Algebra.Ring.Pi import Mathlib.Data.Finsupp.Defs diff --git a/Mathlib/Data/Int/AbsoluteValue.lean b/Mathlib/Data/Int/AbsoluteValue.lean index f061b555444e8..f37f06dffef6a 100644 --- a/Mathlib/Data/Int/AbsoluteValue.lean +++ b/Mathlib/Data/Int/AbsoluteValue.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ import Mathlib.Algebra.GroupWithZero.Action.Units -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.Module.Basic import Mathlib.Algebra.Order.AbsoluteValue import Mathlib.Algebra.Ring.Int.Units import Mathlib.Data.Int.Cast.Lemmas diff --git a/Mathlib/Data/Matrix/Diagonal.lean b/Mathlib/Data/Matrix/Diagonal.lean index d9d2d31a41ec1..6e226513da340 100644 --- a/Mathlib/Data/Matrix/Diagonal.lean +++ b/Mathlib/Data/Matrix/Diagonal.lean @@ -3,7 +3,9 @@ Copyright (c) 2018 Ellen Arlt. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Ellen Arlt, Blair Shi, Sean Leather, Mario Carneiro, Johan Commelin, Lu-Ming Zhang -/ +import Mathlib.Data.Int.Cast.Lemmas import Mathlib.Data.Matrix.Defs +import Mathlib.Data.Nat.Cast.Basic /-! # Diagonal matrices diff --git a/Mathlib/Data/Set/Pointwise/SMul.lean b/Mathlib/Data/Set/Pointwise/SMul.lean index 68200c4096290..1f4a5da3be5cd 100644 --- a/Mathlib/Data/Set/Pointwise/SMul.lean +++ b/Mathlib/Data/Set/Pointwise/SMul.lean @@ -6,6 +6,7 @@ Authors: Johan Commelin, Floris van Doorn import Mathlib.Algebra.Group.Pi.Basic import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.GroupWithZero.Action.Basic +import Mathlib.Algebra.GroupWithZero.Action.Units import Mathlib.Algebra.Module.Defs import Mathlib.Algebra.Ring.Opposite import Mathlib.Algebra.NoZeroSMulDivisors.Defs diff --git a/Mathlib/GroupTheory/ArchimedeanDensely.lean b/Mathlib/GroupTheory/ArchimedeanDensely.lean index 0434db1da4ce6..6e65e049cb576 100644 --- a/Mathlib/GroupTheory/ArchimedeanDensely.lean +++ b/Mathlib/GroupTheory/ArchimedeanDensely.lean @@ -3,12 +3,13 @@ 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.Data.Int.Interval -import Mathlib.GroupTheory.Archimedean import Mathlib.Algebra.Group.Equiv.TypeTags import Mathlib.Algebra.Group.Subgroup.Pointwise +import Mathlib.Algebra.Module.NatInt import Mathlib.Algebra.Order.Group.TypeTags import Mathlib.Algebra.Order.Hom.Monoid +import Mathlib.Data.Int.Interval +import Mathlib.GroupTheory.Archimedean /-! # Archimedean groups are either discrete or densely ordered diff --git a/Mathlib/GroupTheory/Divisible.lean b/Mathlib/GroupTheory/Divisible.lean index 99ead5b297aa7..5543e8f8d13df 100644 --- a/Mathlib/GroupTheory/Divisible.lean +++ b/Mathlib/GroupTheory/Divisible.lean @@ -5,6 +5,7 @@ Authors: Jujian Zhang -/ import Mathlib.Algebra.Group.ULift import Mathlib.Algebra.Group.Subgroup.Pointwise +import Mathlib.Algebra.Module.NatInt import Mathlib.GroupTheory.QuotientGroup.Defs import Mathlib.Tactic.NormNum.Eq diff --git a/Mathlib/GroupTheory/FreeAbelianGroup.lean b/Mathlib/GroupTheory/FreeAbelianGroup.lean index 163b552783eb7..ad652f38aa3fc 100644 --- a/Mathlib/GroupTheory/FreeAbelianGroup.lean +++ b/Mathlib/GroupTheory/FreeAbelianGroup.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ +import Mathlib.Algebra.Module.NatInt import Mathlib.GroupTheory.Abelianization import Mathlib.GroupTheory.FreeGroup.Basic diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index f01189b606a68..b380c1f71c410 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl, Julian Kuelshammer -/ import Mathlib.Algebra.CharP.Defs import Mathlib.Algebra.Group.Subgroup.Finite +import Mathlib.Algebra.Module.NatInt import Mathlib.Algebra.Order.Group.Action import Mathlib.Algebra.Order.Ring.Abs import Mathlib.GroupTheory.Index diff --git a/Mathlib/GroupTheory/QuotientGroup/Basic.lean b/Mathlib/GroupTheory/QuotientGroup/Basic.lean index 4495ebe68b078..3c1822db6828d 100644 --- a/Mathlib/GroupTheory/QuotientGroup/Basic.lean +++ b/Mathlib/GroupTheory/QuotientGroup/Basic.lean @@ -6,6 +6,7 @@ Authors: Kevin Buzzard, Patrick Massot -- This file is to a certain extent based on `quotient_module.lean` by Johannes Hölzl. import Mathlib.Algebra.Group.Subgroup.Pointwise +import Mathlib.Data.Int.Cast.Lemmas import Mathlib.GroupTheory.Congruence.Hom import Mathlib.GroupTheory.Coset.Basic import Mathlib.GroupTheory.QuotientGroup.Defs diff --git a/Mathlib/Order/Filter/Germ/Basic.lean b/Mathlib/Order/Filter/Germ/Basic.lean index d855beedbbf22..21eda99605828 100644 --- a/Mathlib/Order/Filter/Germ/Basic.lean +++ b/Mathlib/Order/Filter/Germ/Basic.lean @@ -3,8 +3,9 @@ Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Abhimanyu Pallavi Sudhir -/ -import Mathlib.Order.Filter.Tendsto import Mathlib.Algebra.Module.Pi +import Mathlib.Data.Int.Cast.Lemmas +import Mathlib.Order.Filter.Tendsto /-! # Germ of a function at a filter From 6095b78b2e3595eb24b62993b94f120a83f7fa91 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Wed, 20 Nov 2024 10:13:28 +0000 Subject: [PATCH 135/829] feat(ModelTheory): Ax-Grothendieck (#6468) Co-authored-by: ChrisHughes24 Co-authored-by: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com> --- Mathlib/FieldTheory/AxGrothendieck.lean | 267 ++++++++++++++++++++---- 1 file changed, 225 insertions(+), 42 deletions(-) diff --git a/Mathlib/FieldTheory/AxGrothendieck.lean b/Mathlib/FieldTheory/AxGrothendieck.lean index 9ddfa55c77de3..b3ce93355bef9 100644 --- a/Mathlib/FieldTheory/AxGrothendieck.lean +++ b/Mathlib/FieldTheory/AxGrothendieck.lean @@ -3,61 +3,244 @@ Copyright (c) 2023 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import Mathlib.Algebra.MvPolynomial.Basic -import Mathlib.Data.Fintype.Card + import Mathlib.RingTheory.Algebraic +import Mathlib.Data.Fintype.Card +import Mathlib.ModelTheory.Algebra.Field.IsAlgClosed +import Mathlib.ModelTheory.Algebra.Ring.Definability +import Mathlib.RingTheory.Polynomial.Basic /-! -# Ax-Grothendieck for algebraic extensions of `ZMod p` +# Ax-Grothendieck -This file proves that if `R` is an algebraic extension of a finite field, -then any injective polynomial map `R^n → R^n` is also surjective. +This file proves that if `K` is an algebraically closed field, +then any injective polynomial map `K^n → K^n` is also surjective. -This proof is required for the true Ax-Grothendieck theorem, which proves the same result -for any algebraically closed field of characteristic zero. +## Main results -## TODO +* `ax_grothendieck_zeroLocus`: If `K` is algebraically closed, `ι` is a finite type, and +`S : Set (ι → K)` is the `zeroLocus` of some ideal of `MvPolynomial ι K`, then any injective +polynomial map `S → S` is also surjective on `S`. +* `ax_grothendieck_univ`: Any injective polynomial map `K^n → K^n` is also surjective if `K` is an +algberaically closed field. +* `ax_grothendieck_of_definable`: Any injective polynomial map `S → S` is also surjective on `S` if +`K` is an algebraically closed field and `S` is a definable subset of `K^n`. +* `ax_grothendieck_of_locally_finite`: any injective polynomial map `R^n → R^n` is also surjective +whenever `R` is an algebraic extension of a finite field. + +## References + +The first order theory of algebraically closed fields, along with the Lefschetz Principle and +the Ax-Grothendieck Theorem were first formalized in Lean 3 by Joseph Hua +[here](https://github.com/Jlh18/ModelTheoryInLean8) with the master's thesis +[here](https://github.com/Jlh18/ModelTheory8Report) -The proof of the theorem for characteristic zero is not in mathlib, but it is at -https://github.com/Jlh18/ModelTheoryInLean8 -/ noncomputable section -open MvPolynomial Finset Function +open MvPolynomial Finset /-- Any injective polynomial map over an algebraic extension of a finite field is surjective. -/ theorem ax_grothendieck_of_locally_finite {ι K R : Type*} [Field K] [Finite K] [CommRing R] - [Finite ι] [Algebra K R] [Algebra.IsAlgebraic K R] (ps : ι → MvPolynomial ι R) - (hinj : Injective fun v i => MvPolynomial.eval v (ps i)) : - Surjective fun v i => MvPolynomial.eval v (ps i) := by + [Finite ι] [Algebra K R] [alg : Algebra.IsAlgebraic K R] (ps : ι → MvPolynomial ι R) + (S : Set (ι → R)) + (hm : S.MapsTo (fun v i => eval v (ps i)) S) + (hinj : S.InjOn (fun v i => eval v (ps i))) : + S.SurjOn (fun v i => eval v (ps i)) S := by + have is_int : ∀ x : R, IsIntegral K x := fun x => isAlgebraic_iff_isIntegral.1 + (alg.isAlgebraic x) + classical + intro v hvS + cases nonempty_fintype ι + /- `s` is the set of all coefficients of the polynomial, as well as all of + the coordinates of `v`, the point I am trying to find the preimage of. -/ + let s : Finset R := + (Finset.biUnion (univ : Finset ι) fun i => (ps i).support.image fun x => coeff x (ps i)) ∪ + (univ : Finset ι).image v + have hv : ∀ i, v i ∈ Algebra.adjoin K (s : Set R) := fun j => + Algebra.subset_adjoin (mem_union_right _ (mem_image.2 ⟨j, mem_univ _, rfl⟩)) + have hs₁ : ∀ (i : ι) (k : ι →₀ ℕ), + k ∈ (ps i).support → coeff k (ps i) ∈ Algebra.adjoin K (s : Set R) := + fun i k hk => Algebra.subset_adjoin + (mem_union_left _ (mem_biUnion.2 ⟨i, mem_univ _, mem_image_of_mem _ hk⟩)) + have := isNoetherian_adjoin_finset s fun x _ => is_int x + have := Module.IsNoetherian.finite K (Algebra.adjoin K (s : Set R)) + have : Finite (Algebra.adjoin K (s : Set R)) := Module.finite_of_finite K + -- The restriction of the polynomial map, `ps`, to the subalgebra generated by `s` + let S' : Set (ι → Algebra.adjoin K (s : Set R)) := + (fun v => Subtype.val ∘ v) ⁻¹' S + let res : S' → S' := fun x => ⟨fun i => + ⟨eval (fun j : ι => (x.1 j : R)) (ps i), eval_mem (hs₁ _) fun i => (x.1 i).2⟩, + hm x.2⟩ + have hres_surj : Function.Surjective res := by + rw [← Finite.injective_iff_surjective] + intro x y hxy + ext i + simp only [Subtype.ext_iff, funext_iff] at hxy + exact congr_fun (hinj x.2 y.2 (funext hxy)) i + rcases hres_surj ⟨fun i => ⟨v i, hv i⟩, hvS⟩ with ⟨⟨w, hwS'⟩, hw⟩ + refine ⟨fun i => w i, hwS', ?_⟩ + simpa [Subtype.ext_iff, funext_iff] using hw + +end + +namespace FirstOrder + +open MvPolynomial FreeCommRing Language Field Ring BoundedFormula + +variable {ι α : Type*} [Finite α] {K : Type*} [Field K] [CompatibleRing K] + +/-- The collection of first order formulas corresponding to the Ax-Grothendieck theorem. -/ +noncomputable def genericPolyMapSurjOnOfInjOn [Fintype ι] + (φ : ring.Formula (α ⊕ ι)) + (mons : ι → Finset (ι →₀ ℕ)) : Language.ring.Sentence := + let l1 : ι → Language.ring.Formula ((Σ i : ι, mons i) ⊕ (Fin 2 × ι)) := + fun i => + (termOfFreeCommRing (genericPolyMap mons i)).relabel + (Sum.inl ∘ Sum.map id (fun i => (0, i))) + =' (termOfFreeCommRing (genericPolyMap mons i)).relabel + (Sum.inl ∘ Sum.map id (fun i => (1, i))) + -- p(x) = p(y) as a formula + let f1 : Language.ring.Formula ((Σ i : ι, mons i) ⊕ (Fin 2 × ι)) := + iInf Finset.univ l1 + let l2 : ι → Language.ring.Formula ((Σ i : ι, mons i) ⊕ (Fin 2 × ι)) := + fun i => .var (Sum.inl (Sum.inr (0, i))) =' .var (Sum.inl (Sum.inr (1, i))) + -- x = y as a formula + let f2 : Language.ring.Formula ((Σ i : ι, mons i) ⊕ (Fin 2 × ι)) := + iInf Finset.univ l2 + let injOn : Language.ring.Formula (α ⊕ Σ i : ι, mons i) := + Formula.iAlls (γ := Fin 2 × ι) id + (φ.relabel (Sum.map Sum.inl (fun i => (0, i))) ⟹ + φ.relabel (Sum.map Sum.inl (fun i => (1, i))) ⟹ + (f1.imp f2).relabel (fun x => (Equiv.sumAssoc _ _ _).symm (Sum.inr x))) + let l3 : ι → Language.ring.Formula ((Σ i : ι, mons i) ⊕ (Fin 2 × ι)) := + fun i => (termOfFreeCommRing (genericPolyMap mons i)).relabel + (Sum.inl ∘ Sum.map id (fun i => (0, i))) =' + .var (Sum.inl (Sum.inr (1, i))) + let f3 : Language.ring.Formula ((Σ i : ι, mons i) ⊕ (Fin 2 × ι)) := + iInf Finset.univ l3 + let surjOn : Language.ring.Formula (α ⊕ Σ i : ι, mons i) := + Formula.iAlls (γ := ι) id + (Formula.imp (φ.relabel (Sum.map Sum.inl id)) <| + Formula.iExs (γ := ι) + (fun (i : (α ⊕ (Σ i : ι, mons i)) ⊕ (Fin 2 × ι)) => + show ((α ⊕ (Σ i : ι, mons i)) ⊕ ι) ⊕ ι + from Sum.elim (Sum.inl ∘ Sum.inl) + (fun i => if i.1 = 0 then Sum.inr i.2 else (Sum.inl (Sum.inr i.2))) i) + ((φ.relabel (Sum.map Sum.inl (fun i => (0, i)))) ⊓ + (f3.relabel (fun x => (Equiv.sumAssoc _ _ _).symm (Sum.inr x))))) + let mapsTo : Language.ring.Formula (α ⊕ Σ i : ι, mons i) := + Formula.iAlls (γ := ι) id + (Formula.imp (φ.relabel (Sum.map Sum.inl id)) + (φ.subst <| Sum.elim + (fun a => .var (Sum.inl (Sum.inl a))) + (fun i => (termOfFreeCommRing (genericPolyMap mons i)).relabel + (fun i => (Equiv.sumAssoc _ _ _).symm (Sum.inr i))))) + Formula.iAlls (γ := α ⊕ Σ i : ι, mons i) Sum.inr (mapsTo ⟹ injOn ⟹ surjOn) + +theorem realize_genericPolyMapSurjOnOfInjOn + [Fintype ι] (φ : ring.Formula (α ⊕ ι)) (mons : ι → Finset (ι →₀ ℕ)) : + (K ⊨ genericPolyMapSurjOnOfInjOn φ mons) ↔ + ∀ (v : α → K) (p : { p : ι → MvPolynomial ι K // (∀ i, (p i).support ⊆ mons i) }), + let f : (ι → K) → (ι → K) := fun v i => eval v (p.1 i) + let S : Set (ι → K) := fun x => φ.Realize (Sum.elim v x) + S.MapsTo f S → S.InjOn f → S.SurjOn f S := by classical - intro v - cases nonempty_fintype ι - /- `s` is the set of all coefficients of the polynomial, as well as all of - the coordinates of `v`, the point I am trying to find the preimage of. -/ - let s : Finset R := - (Finset.biUnion (univ : Finset ι) fun i => (ps i).support.image fun x => coeff x (ps i)) ∪ - (univ : Finset ι).image v - have hv : ∀ i, v i ∈ Algebra.adjoin K (s : Set R) := fun j => - Algebra.subset_adjoin (mem_union_right _ (mem_image.2 ⟨j, mem_univ _, rfl⟩)) - have hs₁ : ∀ (i : ι) (k : ι →₀ ℕ), - k ∈ (ps i).support → coeff k (ps i) ∈ Algebra.adjoin K (s : Set R) := - fun i k hk => Algebra.subset_adjoin - (mem_union_left _ (mem_biUnion.2 ⟨i, mem_univ _, mem_image_of_mem _ hk⟩)) - letI := isNoetherian_adjoin_finset s fun x _ => Algebra.IsIntegral.isIntegral (R := K) x - letI := Module.IsNoetherian.finite K (Algebra.adjoin K (s : Set R)) - letI : Finite (Algebra.adjoin K (s : Set R)) := Module.finite_of_finite K - -- The restriction of the polynomial map, `ps`, to the subalgebra generated by `s` - let res : (ι → Algebra.adjoin K (s : Set R)) → ι → Algebra.adjoin K (s : Set R) := fun x i => - ⟨eval (fun j : ι => (x j : R)) (ps i), eval_mem (hs₁ _) fun i => (x i).2⟩ - have hres_inj : Injective res := by - intro x y hxy - ext i - simp only [Subtype.ext_iff, funext_iff] at hxy - exact congr_fun (hinj (funext hxy)) i - have hres_surj : Surjective res := Finite.injective_iff_surjective.1 hres_inj - cases' hres_surj fun i => ⟨v i, hv i⟩ with w hw - use fun i => w i - simpa only [Subtype.ext_iff, funext_iff] using hw + have injOnAlt : ∀ {S : Set (ι → K)} (f : (ι → K) → (ι → K)), + S.InjOn f ↔ ∀ x y, x ∈ S → y ∈ S → f x = f y → x = y := by + simp [Set.InjOn]; tauto + simp only [Sentence.Realize, Formula.Realize, genericPolyMapSurjOnOfInjOn, Formula.relabel, + Function.comp_def, Sum.map, id_eq, Equiv.sumAssoc, Equiv.coe_fn_symm_mk, Sum.elim_inr, + realize_iAlls, realize_imp, realize_relabel, Fin.natAdd_zero, realize_subst, realize_iInf, + Finset.mem_univ, realize_bdEqual, Term.realize_relabel, true_imp_iff, + Equiv.forall_congr_left (Equiv.curry (Fin 2) ι K), Equiv.curry_symm_apply, Function.uncurry, + Fin.forall_fin_succ_pi, Fin.forall_fin_zero_pi, realize_iExs, realize_inf, Sum.forall_sum, + Set.MapsTo, Set.mem_def, injOnAlt, funext_iff, Set.SurjOn, Set.image, setOf, + Set.subset_def, Equiv.forall_congr_left (mvPolynomialSupportLEEquiv mons)] + simp +singlePass only [← Sum.elim_comp_inl_inr] + simp [Set.mem_def, Function.comp_def] + +theorem ACF_models_genericPolyMapSurjOnOfInjOn_of_prime [Fintype ι] + {p : ℕ} (hp : p.Prime) (φ : ring.Formula (α ⊕ ι)) (mons : ι → Finset (ι →₀ ℕ)) : + Theory.ACF p ⊨ᵇ genericPolyMapSurjOnOfInjOn φ mons := by + classical + have : Fact p.Prime := ⟨hp⟩ + letI := compatibleRingOfRing (AlgebraicClosure (ZMod p)) + have : CharP (AlgebraicClosure (ZMod p)) p := + charP_of_injective_algebraMap + (RingHom.injective (algebraMap (ZMod p) (AlgebraicClosure (ZMod p)))) p + rw [← (ACF_isComplete (Or.inl hp)).realize_sentence_iff _ + (AlgebraicClosure (ZMod p)), realize_genericPolyMapSurjOnOfInjOn] + rintro v ⟨f, _⟩ + exact ax_grothendieck_of_locally_finite (K := ZMod p) (ι := ι) f _ + +theorem ACF_models_genericPolyMapSurjOnOfInjOn_of_prime_or_zero + [Fintype ι] {p : ℕ} (hp : p.Prime ∨ p = 0) + (φ : ring.Formula (α ⊕ ι)) (mons : ι → Finset (ι →₀ ℕ)) : + Theory.ACF p ⊨ᵇ genericPolyMapSurjOnOfInjOn φ mons := by + rcases hp with hp | rfl + · exact ACF_models_genericPolyMapSurjOnOfInjOn_of_prime hp φ mons + · rw [ACF_zero_realize_iff_infinite_ACF_prime_realize] + convert Set.infinite_univ (α := Nat.Primes) + rw [Set.eq_univ_iff_forall] + intro ⟨p, hp⟩ + exact ACF_models_genericPolyMapSurjOnOfInjOn_of_prime hp φ mons + +end FirstOrder + +open FirstOrder Language Field Ring MvPolynomial + +variable {K ι : Type*} [Field K] [IsAlgClosed K] [Finite ι] + +/-- A slight generalization of the **Ax-Grothendieck** theorem + +If `K` is an algebraically closed field, `ι` is a finite type, and `S` is a definable subset of +`ι → K`, then any injective polynomial map `S → S` is also surjective on `S`. -/ +theorem ax_grothendieck_of_definable [CompatibleRing K] {c : Set K} + (S : Set (ι → K)) (hS : c.Definable Language.ring S) + (ps : ι → MvPolynomial ι K) : + S.MapsTo (fun v i => eval v (ps i)) S → + S.InjOn (fun v i => eval v (ps i)) → + S.SurjOn (fun v i => eval v (ps i)) S := by + letI := Fintype.ofFinite ι + let p : ℕ := ringChar K + have : CharP K p := ⟨ringChar.spec K⟩ + rw [Set.definable_iff_finitely_definable] at hS + rcases hS with ⟨c, _, hS⟩ + rw [Set.definable_iff_exists_formula_sum] at hS + rcases hS with ⟨φ, hφ⟩ + rw [hφ] + have := ACF_models_genericPolyMapSurjOnOfInjOn_of_prime_or_zero + (CharP.char_is_prime_or_zero K p) φ (fun i => (ps i).support) + rw [← (ACF_isComplete (CharP.char_is_prime_or_zero K p)).realize_sentence_iff _ K, + realize_genericPolyMapSurjOnOfInjOn] at this + exact this Subtype.val ⟨ps, fun i => Set.Subset.refl _⟩ + +/-- The **Ax-Grothendieck** theorem + +If `K` is an algebraically closed field, and `S : Set (ι → K)` is the `zeroLocus` of an ideal +of the multivariable polynomial ring, then any injective polynomial map `S → S` is also +surjective on `S`. -/ +theorem ax_grothendieck_zeroLocus + (I : Ideal (MvPolynomial ι K)) + (p : ι → MvPolynomial ι K) : + let S := zeroLocus I + S.MapsTo (fun v i => eval v (p i)) S → + S.InjOn (fun v i => eval v (p i)) → + S.SurjOn (fun v i => eval v (p i)) S := by + letI := compatibleRingOfRing K + intro S + obtain ⟨s, rfl⟩ : I.FG := IsNoetherian.noetherian I + exact ax_grothendieck_of_definable S (mvPolynomial_zeroLocus_definable s) p + +/-- A special case of the **Ax-Grothendieck** theorem + +Any injective polynomial map `K^n → K^n` is also surjective if `K` is an +algberaically closed field. -/ +theorem ax_grothendieck_univ (p : ι → MvPolynomial ι K) : + (fun v i => eval v (p i)).Injective → + (fun v i => eval v (p i)).Surjective := by + simpa [Set.injective_iff_injOn_univ, Set.surjective_iff_surjOn_univ] using + ax_grothendieck_zeroLocus 0 p From ad04df3876a303393c2a7a9e067bcecefce8a678 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Wed, 20 Nov 2024 10:44:46 +0000 Subject: [PATCH 136/829] feat: bounds for modular forms of non-positive weights (#18192) We use the Maximum modulus principle to show that modular forms of non-positive weights are bounded by its value of some element of the upper half plane imaginary part at least 1/2. Co-authored-by: Chris Birkbeck Co-authored-by: David Loeffler --- Mathlib.lean | 1 + .../Complex/UpperHalfPlane/Basic.lean | 158 +++++++++--------- Mathlib/NumberTheory/Modular.lean | 34 +++- .../ModularForms/EisensteinSeries/Defs.lean | 4 +- .../NumberTheory/ModularForms/LevelOne.lean | 38 +++++ .../ModularForms/SlashActions.lean | 6 +- 6 files changed, 153 insertions(+), 88 deletions(-) create mode 100644 Mathlib/NumberTheory/ModularForms/LevelOne.lean diff --git a/Mathlib.lean b/Mathlib.lean index 797b5848ce70d..f0d237abd75a7 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3788,6 +3788,7 @@ import Mathlib.NumberTheory.ModularForms.JacobiTheta.Bounds import Mathlib.NumberTheory.ModularForms.JacobiTheta.Manifold import Mathlib.NumberTheory.ModularForms.JacobiTheta.OneVariable import Mathlib.NumberTheory.ModularForms.JacobiTheta.TwoVariable +import Mathlib.NumberTheory.ModularForms.LevelOne import Mathlib.NumberTheory.ModularForms.QExpansion import Mathlib.NumberTheory.ModularForms.SlashActions import Mathlib.NumberTheory.ModularForms.SlashInvariantForms diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean index 52bf76205cdd1..0ac5af563ff8a 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean @@ -268,50 +268,9 @@ instance : MulAction GL(2, ℝ)⁺ ℍ where simp [num, denom] mul_smul := mul_smul' -section ModularScalarTowers - instance SLAction {R : Type*} [CommRing R] [Algebra R ℝ] : MulAction SL(2, R) ℍ := MulAction.compHom ℍ <| SpecialLinearGroup.toGLPos.comp <| map (algebraMap R ℝ) -namespace ModularGroup - -variable (Γ : Subgroup (SpecialLinearGroup (Fin 2) ℤ)) - -/-- Canonical embedding of `SL(2, ℤ)` into `GL(2, ℝ)⁺`. -/ -@[coe] -def coe' : SL(2, ℤ) → GL(2, ℝ)⁺ := fun g => ((g : SL(2, ℝ)) : GL(2, ℝ)⁺) - -instance : Coe SL(2, ℤ) GL(2, ℝ)⁺ := - ⟨coe'⟩ - -@[simp] -theorem coe'_apply_complex {g : SL(2, ℤ)} {i j : Fin 2} : - (Units.val <| Subtype.val <| coe' g) i j = (Subtype.val g i j : ℂ) := - rfl - -@[simp] -theorem det_coe' {g : SL(2, ℤ)} : det (Units.val <| Subtype.val <| coe' g) = 1 := by - simp only [SpecialLinearGroup.coe_GLPos_coe_GL_coe_matrix, SpecialLinearGroup.det_coe, coe'] - -lemma coe_one : UpperHalfPlane.ModularGroup.coe' 1 = 1 := by - simp only [coe', _root_.map_one] - -instance SLOnGLPos : SMul SL(2, ℤ) GL(2, ℝ)⁺ := - ⟨fun s g => s * g⟩ - -theorem SLOnGLPos_smul_apply (s : SL(2, ℤ)) (g : GL(2, ℝ)⁺) (z : ℍ) : - (s • g) • z = ((s : GL(2, ℝ)⁺) * g) • z := - rfl - -instance SL_to_GL_tower : IsScalarTower SL(2, ℤ) GL(2, ℝ)⁺ ℍ where - smul_assoc s g z := by - simp only [SLOnGLPos_smul_apply] - apply mul_smul' - -end ModularGroup - -end ModularScalarTowers - -- Porting note: in the statement, we used to have coercions `↑· : ℝ` -- rather than `algebraMap R ℝ ·`. theorem specialLinearGroup_apply {R : Type*} [CommRing R] [Algebra R ℝ] (g : SL(2, R)) (z : ℍ) : @@ -322,23 +281,23 @@ theorem specialLinearGroup_apply {R : Type*} [CommRing R] [Algebra R ℝ] (g : S (g • z).property := rfl +variable (g : GL(2, ℝ)⁺) (z : ℍ) + @[simp] -theorem coe_smul (g : GL(2, ℝ)⁺) (z : ℍ) : ↑(g • z) = num g z / denom g z := +theorem coe_smul : ↑(g • z) = num g z / denom g z := rfl @[simp] -theorem re_smul (g : GL(2, ℝ)⁺) (z : ℍ) : (g • z).re = (num g z / denom g z).re := +theorem re_smul : (g • z).re = (num g z / denom g z).re := rfl -theorem im_smul (g : GL(2, ℝ)⁺) (z : ℍ) : (g • z).im = (num g z / denom g z).im := +theorem im_smul : (g • z).im = (num g z / denom g z).im := rfl -theorem im_smul_eq_div_normSq (g : GL(2, ℝ)⁺) (z : ℍ) : - (g • z).im = det ↑ₘg * z.im / Complex.normSq (denom g z) := +theorem im_smul_eq_div_normSq : (g • z).im = det ↑ₘg * z.im / Complex.normSq (denom g z) := smulAux'_im g z -theorem c_mul_im_sq_le_normSq_denom (z : ℍ) (g : SL(2, ℝ)) : - (g 1 0 * z.im) ^ 2 ≤ Complex.normSq (denom g z) := by +theorem c_mul_im_sq_le_normSq_denom : (g 1 0 * z.im) ^ 2 ≤ Complex.normSq (denom g z) := by set c := g 1 0 set d := g 1 1 calc @@ -346,40 +305,15 @@ theorem c_mul_im_sq_le_normSq_denom (z : ℍ) (g : SL(2, ℝ)) : _ = Complex.normSq (denom g z) := by dsimp [c, d, denom, Complex.normSq]; ring @[simp] -theorem neg_smul (g : GL(2, ℝ)⁺) (z : ℍ) : -g • z = g • z := by +theorem neg_smul : -g • z = g • z := by ext1 change _ / _ = _ / _ field_simp [denom_ne_zero] simp only [num, denom, Complex.ofReal_neg, neg_mul, GLPos.coe_neg_GL, Units.val_neg, neg_apply] ring_nf -section SLModularAction - -namespace ModularGroup - -variable (g : SL(2, ℤ)) (z : ℍ) (Γ : Subgroup SL(2, ℤ)) - -@[simp] -theorem sl_moeb (A : SL(2, ℤ)) (z : ℍ) : A • z = (A : GL(2, ℝ)⁺) • z := - rfl - -@[simp high] -theorem SL_neg_smul (g : SL(2, ℤ)) (z : ℍ) : -g • z = g • z := by - simp only [coe_GLPos_neg, sl_moeb, coe_int_neg, neg_smul, coe'] - -nonrec theorem im_smul_eq_div_normSq : - (g • z).im = z.im / Complex.normSq (denom g z) := by - convert im_smul_eq_div_normSq g z - simp only [GeneralLinearGroup.val_det_apply, coe_GLPos_coe_GL_coe_matrix, - Int.coe_castRingHom, (g : SL(2, ℝ)).prop, one_mul, coe'] - -theorem denom_apply (g : SL(2, ℤ)) (z : ℍ) : - denom g z = (↑g : Matrix (Fin 2) (Fin 2) ℤ) 1 0 * z + (↑g : Matrix (Fin 2) (Fin 2) ℤ) 1 1 := by - simp [denom, coe'] - -end ModularGroup - -end SLModularAction +lemma denom_one : denom 1 z = 1 := by + simp [denom] section PosRealAction @@ -475,3 +409,75 @@ theorem exists_SL2_smul_eq_of_apply_zero_one_ne_zero (g : SL(2, ℝ)) (hc : g 1 linear_combination (-(z * (c : ℂ) ^ 2) - c * d) * h end UpperHalfPlane + +namespace ModularGroup -- results specific to `SL(2, ℤ)` + +section ModularScalarTowers + +/-- Canonical embedding of `SL(2, ℤ)` into `GL(2, ℝ)⁺`. -/ +@[coe] +def coe (g : SL(2, ℤ)) : GL(2, ℝ)⁺ := ((g : SL(2, ℝ)) : GL(2, ℝ)⁺) + +@[deprecated (since := "2024-11-19")] noncomputable alias coe' := coe + +instance : Coe SL(2, ℤ) GL(2, ℝ)⁺ := + ⟨coe⟩ + +@[simp] +theorem coe_apply_complex {g : SL(2, ℤ)} {i j : Fin 2} : + (Units.val <| Subtype.val <| coe g) i j = (Subtype.val g i j : ℂ) := + rfl + +@[deprecated (since := "2024-11-19")] alias coe'_apply_complex := coe_apply_complex + +@[simp] +theorem det_coe {g : SL(2, ℤ)} : det (Units.val <| Subtype.val <| coe g) = 1 := by + simp only [SpecialLinearGroup.coe_GLPos_coe_GL_coe_matrix, SpecialLinearGroup.det_coe, coe] + +@[deprecated (since := "2024-11-19")] alias det_coe' := det_coe + +lemma coe_one : coe 1 = 1 := by + simp only [coe, _root_.map_one] + +instance SLOnGLPos : SMul SL(2, ℤ) GL(2, ℝ)⁺ := + ⟨fun s g => s * g⟩ + +theorem SLOnGLPos_smul_apply (s : SL(2, ℤ)) (g : GL(2, ℝ)⁺) (z : ℍ) : + (s • g) • z = ((s : GL(2, ℝ)⁺) * g) • z := + rfl + +instance SL_to_GL_tower : IsScalarTower SL(2, ℤ) GL(2, ℝ)⁺ ℍ where + smul_assoc s g z := by + simp only [SLOnGLPos_smul_apply] + apply mul_smul' + +end ModularScalarTowers + +section SLModularAction + +variable (g : SL(2, ℤ)) (z : ℍ) + +@[simp] +theorem sl_moeb (A : SL(2, ℤ)) (z : ℍ) : A • z = (A : GL(2, ℝ)⁺) • z := + rfl + +@[simp high] +theorem SL_neg_smul (g : SL(2, ℤ)) (z : ℍ) : -g • z = g • z := by + simp only [coe_GLPos_neg, sl_moeb, coe_int_neg, neg_smul, coe] + +theorem im_smul_eq_div_normSq : (g • z).im = z.im / Complex.normSq (denom g z) := by + simpa only [coe, coe_GLPos_coe_GL_coe_matrix, (g : SL(2, ℝ)).prop, one_mul] using + z.im_smul_eq_div_normSq g + +theorem denom_apply (g : SL(2, ℤ)) (z : ℍ) : + denom g z = g 1 0 * z + g 1 1 := by + simp [denom, coe] + +@[simp] +lemma denom_S (z : ℍ) : denom S z = z := by + simp only [S, denom_apply, of_apply, cons_val', cons_val_zero, empty_val', cons_val_fin_one, + cons_val_one, head_fin_const, Int.cast_one, one_mul, head_cons, Int.cast_zero, add_zero] + +end SLModularAction + +end ModularGroup diff --git a/Mathlib/NumberTheory/Modular.lean b/Mathlib/NumberTheory/Modular.lean index 8f86113a1bed0..f4c849d0063d4 100644 --- a/Mathlib/NumberTheory/Modular.lean +++ b/Mathlib/NumberTheory/Modular.lean @@ -239,7 +239,7 @@ theorem smul_eq_lcRow0_add {p : Fin 2 → ℤ} (hp : IsCoprime (p 0) (p 1)) (hg (p 1 : ℂ) * z - p 0 = (p 1 * z - p 0) * ↑(Matrix.det (↑g : Matrix (Fin 2) (Fin 2) ℤ)))] rw [← hg, det_fin_two] simp only [Int.coe_castRingHom, coe_matrix_coe, Int.cast_mul, ofReal_intCast, map_apply, denom, - Int.cast_sub, coe_GLPos_coe_GL_coe_matrix, coe'_apply_complex] + Int.cast_sub, coe_GLPos_coe_GL_coe_matrix, coe_apply_complex] ring theorem tendsto_abs_re_smul {p : Fin 2 → ℤ} (hp : IsCoprime (p 0) (p 1)) : @@ -384,6 +384,11 @@ theorem three_lt_four_mul_im_sq_of_mem_fdo (h : z ∈ 𝒟ᵒ) : 3 < 4 * z.im ^ have := h.2 cases abs_cases z.re <;> nlinarith +/-- non-strict variant of `ModularGroup.three_le_four_mul_im_sq_of_mem_fdo` -/ +theorem three_le_four_mul_im_sq_of_mem_fd {τ : ℍ} (h : τ ∈ 𝒟) : 3 ≤ 4 * τ.im ^ 2 := by + have : 1 ≤ τ.re * τ.re + τ.im * τ.im := by simpa [Complex.normSq_apply] using h.1 + cases abs_cases τ.re <;> nlinarith [h.2] + /-- If `z ∈ 𝒟ᵒ`, and `n : ℤ`, then `|z + n| > 1`. -/ theorem one_lt_normSq_T_zpow_smul (hz : z ∈ 𝒟ᵒ) (n : ℤ) : 1 < normSq (T ^ n • z : ℍ) := by have hz₁ : 1 < z.re * z.re + z.im * z.im := hz.1 @@ -455,20 +460,18 @@ theorem abs_c_le_one (hz : z ∈ 𝒟ᵒ) (hg : g • z ∈ 𝒟ᵒ) : |g 1 0| intro hc replace hc : 0 < c ^ 4 := by change 0 < c ^ (2 * 2); rw [pow_mul]; apply sq_pos_of_pos (sq_pos_of_ne_zero hc) - have h₁ := - mul_lt_mul_of_pos_right + have h₁ := mul_lt_mul_of_pos_right (mul_lt_mul'' (three_lt_four_mul_im_sq_of_mem_fdo hg) (three_lt_four_mul_im_sq_of_mem_fdo hz) - (by linarith) (by linarith)) - hc + (by norm_num) (by norm_num)) hc have h₂ : (c * z.im) ^ 4 / normSq (denom (↑g) z) ^ 2 ≤ 1 := div_le_one_of_le₀ - (pow_four_le_pow_two_of_pow_two_le (UpperHalfPlane.c_mul_im_sq_le_normSq_denom z g)) + (pow_four_le_pow_two_of_pow_two_le (z.c_mul_im_sq_le_normSq_denom g)) (sq_nonneg _) let nsq := normSq (denom g z) calc 9 * c ^ 4 < c ^ 4 * z.im ^ 2 * (g • z).im ^ 2 * 16 := by linarith _ = c ^ 4 * z.im ^ 4 / nsq ^ 2 * 16 := by - rw [ModularGroup.im_smul_eq_div_normSq, div_pow] + rw [im_smul_eq_div_normSq, div_pow] ring _ ≤ 16 := by rw [← mul_pow]; linarith @@ -504,4 +507,21 @@ end UniqueRepresentative end FundamentalDomain +lemma exists_one_half_le_im_smul (τ : ℍ) : ∃ γ : SL(2, ℤ), 1 / 2 ≤ im (γ • τ) := by + obtain ⟨γ, hγ⟩ := exists_smul_mem_fd τ + use γ + nlinarith [three_le_four_mul_im_sq_of_mem_fd hγ, im_pos (γ • τ)] + +/-- For every `τ : ℍ` there is some `γ ∈ SL(2, ℤ)` that sends it to an element whose +imaginary part is at least `1/2` and such that `denom γ τ` has norm at most 1. -/ +lemma exists_one_half_le_im_smul_and_norm_denom_le (τ : ℍ) : + ∃ γ : SL(2, ℤ), 1 / 2 ≤ im (γ • τ) ∧ ‖denom γ τ‖ ≤ 1 := by + rcases le_total (1 / 2) τ.im with h | h + · exact ⟨1, (one_smul SL(2, ℤ) τ).symm ▸ h, by simp only [coe_one, denom_one, norm_one, le_refl]⟩ + · refine (exists_one_half_le_im_smul τ).imp (fun γ hγ ↦ ⟨hγ, ?_⟩) + have h1 : τ.im ≤ (γ • τ).im := h.trans hγ + rw [im_smul_eq_div_normSq, le_div_iff₀ (normSq_denom_pos (↑γ) τ), normSq_eq_norm_sq] at h1 + simpa only [norm_eq_abs, sq_le_one_iff_abs_le_one, Complex.abs_abs] using + (mul_le_iff_le_one_right τ.2).mp h1 + end ModularGroup diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Defs.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Defs.lean index 533e632c67595..890e52757badd 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Defs.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Defs.lean @@ -109,10 +109,10 @@ def _root_.eisensteinSeries (k : ℤ) (z : ℍ) : ℂ := ∑' x : gammaSet N a, lemma eisensteinSeries_slash_apply (k : ℤ) (γ : SL(2, ℤ)) : eisensteinSeries a k ∣[k] γ = eisensteinSeries (a ᵥ* γ) k := by ext1 z - simp_rw [SL_slash, slash_def, slash, ModularGroup.det_coe', ofReal_one, one_zpow, mul_one, + simp_rw [SL_slash, slash_def, slash, ModularGroup.det_coe, ofReal_one, one_zpow, mul_one, zpow_neg, mul_inv_eq_iff_eq_mul₀ (zpow_ne_zero _ <| z.denom_ne_zero _), mul_comm, eisensteinSeries, ← ModularGroup.sl_moeb, eisSummand_SL2_apply, tsum_mul_left] - erw [(gammaSetEquiv a γ).tsum_eq (eisSummand k · z)] + exact congr_arg (_ * ·) <| (gammaSetEquiv a γ).tsum_eq (eisSummand k · z) /-- The SlashInvariantForm defined by an Eisenstein series of weight `k : ℤ`, level `Γ(N)`, and congruence condition given by `a : Fin 2 → ZMod N`. -/ diff --git a/Mathlib/NumberTheory/ModularForms/LevelOne.lean b/Mathlib/NumberTheory/ModularForms/LevelOne.lean new file mode 100644 index 0000000000000..20ba662c6cc3b --- /dev/null +++ b/Mathlib/NumberTheory/ModularForms/LevelOne.lean @@ -0,0 +1,38 @@ +/- +Copyright (c) 2024 Chris Birkbeck. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Birkbeck +-/ +import Mathlib.NumberTheory.Modular +import Mathlib.NumberTheory.ModularForms.SlashInvariantForms +/-! +# Level one modular forms + +This file contains results specific to modular forms of level one, ie. modular forms for `SL(2, ℤ)`. + +TODO: Add finite-dimensionality of these spaces of modular forms. + +-/ + +open UpperHalfPlane ModularGroup SlashInvariantForm ModularForm Complex MatrixGroups + +lemma SlashInvariantForm.exists_one_half_le_im_and_norm_le {k : ℤ} (hk : k ≤ 0) {F : Type*} + [FunLike F ℍ ℂ] [SlashInvariantFormClass F ⊤ k] (f : F) (τ : ℍ) : + ∃ ξ : ℍ, 1/2 ≤ ξ.im ∧ ‖f τ‖ ≤ ‖f ξ‖ := + let ⟨γ, hγ, hdenom⟩ := exists_one_half_le_im_smul_and_norm_denom_le τ + ⟨γ • τ, hγ, by simpa only [slash_action_eqn'' _ (show γ ∈ ⊤ by tauto), norm_mul, norm_zpow] + using le_mul_of_one_le_left (norm_nonneg _) <| + one_le_zpow_of_nonpos₀ (norm_pos_iff.2 (denom_ne_zero _ _)) hdenom hk⟩ + +/-- If a constant function is modular of weight `k`, then either `k = 0`, or the constant is `0`. -/ +lemma SlashInvariantForm.wt_eq_zero_of_eq_const + {F : Type*} [FunLike F ℍ ℂ] (k : ℤ) [SlashInvariantFormClass F ⊤ k] + {f : F} {c : ℂ} (hf : ∀ τ, f τ = c) : k = 0 ∨ c = 0 := by + have hI := slash_action_eqn'' f (by tauto : ModularGroup.S ∈ ⊤) I + have h2I2 := slash_action_eqn'' f (by tauto : ModularGroup.S ∈ ⊤) ⟨2 * Complex.I, by simp⟩ + simp only [sl_moeb, hf, denom_S, coe_mk_subtype] at hI h2I2 + nth_rw 1 [h2I2] at hI + simp only [mul_zpow, coe_I, mul_eq_mul_right_iff, mul_left_eq_self₀] at hI + refine hI.imp_left (Or.casesOn · (fun H ↦ ?_) (False.elim ∘ zpow_ne_zero k I_ne_zero)) + rwa [← Complex.ofReal_ofNat, ← ofReal_zpow, ← ofReal_one, ofReal_inj, + zpow_eq_one_iff_right₀ (by norm_num) (by norm_num)] at H diff --git a/Mathlib/NumberTheory/ModularForms/SlashActions.lean b/Mathlib/NumberTheory/ModularForms/SlashActions.lean index 4b39fc4a7a847..bc52891b5dffc 100644 --- a/Mathlib/NumberTheory/ModularForms/SlashActions.lean +++ b/Mathlib/NumberTheory/ModularForms/SlashActions.lean @@ -139,7 +139,7 @@ theorem SL_slash (γ : SL(2, ℤ)) : f ∣[k] γ = f ∣[k] (γ : GL(2, ℝ)⁺) theorem is_invariant_const (A : SL(2, ℤ)) (x : ℂ) : Function.const ℍ x ∣[(0 : ℤ)] A = Function.const ℍ x := by funext - simp only [SL_slash, slash_def, slash, Function.const_apply, det_coe', ofReal_one, zero_sub, + simp only [SL_slash, slash_def, slash, Function.const_apply, det_coe, ofReal_one, zero_sub, zpow_neg, zpow_one, inv_one, mul_one, neg_zero, zpow_zero] /-- The constant function 1 is invariant under any element of `SL(2, ℤ)`. -/ @@ -155,7 +155,7 @@ theorem slash_action_eq'_iff (k : ℤ) (f : ℍ → ℂ) (γ : SL(2, ℤ)) (z : simp only [SL_slash, slash_def, ModularForm.slash] convert inv_mul_eq_iff_eq_mul₀ (G₀ := ℂ) _ using 2 · rw [mul_comm] - simp only [denom, zpow_neg, det_coe', ofReal_one, one_zpow, mul_one, + simp only [denom, zpow_neg, det_coe, ofReal_one, one_zpow, mul_one, sl_moeb] rfl · convert zpow_ne_zero k (denom_ne_zero γ z) @@ -184,7 +184,7 @@ theorem mul_slash_SL2 (k1 k2 : ℤ) (A : SL(2, ℤ)) (f g : ℍ → ℂ) : (f * g) ∣[k1 + k2] (A : GL(2, ℝ)⁺) = ((↑ₘA).det : ℝ) • f ∣[k1] A * g ∣[k2] A := by apply mul_slash - _ = (1 : ℝ) • f ∣[k1] A * g ∣[k2] A := by rw [det_coe'] + _ = (1 : ℝ) • f ∣[k1] A * g ∣[k2] A := by rw [det_coe] _ = f ∣[k1] A * g ∣[k2] A := by rw [one_smul] end From e69c3062daeb950c9dee862d68daa01c08283884 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 20 Nov 2024 10:54:05 +0000 Subject: [PATCH 137/829] chore: deprecate `LinearOrderedCommGroupWithZero` lemmas (#19197) Deprecate almost all lemmas about `LinearOrderedCommGroupWithZero` and remove three months old deprecations. --- Mathlib/Algebra/Order/Archimedean/Basic.lean | 1 + .../Order/GroupWithZero/Canonical.lean | 99 ++++++------------- .../Order/GroupWithZero/Unbundled.lean | 1 - .../Order/GroupWithZero/Unbundled/Lemmas.lean | 32 +++++- Mathlib/Algebra/Order/Hom/Monoid.lean | 1 + Mathlib/RingTheory/LaurentSeries.lean | 3 +- Mathlib/RingTheory/Perfection.lean | 2 +- Mathlib/RingTheory/Valuation/Integral.lean | 2 +- .../Valuation/ValuationSubring.lean | 3 +- .../Algebra/Valued/ValuationTopology.lean | 3 +- .../Topology/Algebra/WithZeroTopology.lean | 2 +- 11 files changed, 70 insertions(+), 79 deletions(-) diff --git a/Mathlib/Algebra/Order/Archimedean/Basic.lean b/Mathlib/Algebra/Order/Archimedean/Basic.lean index 2f526eb1aa929..7756989822b78 100644 --- a/Mathlib/Algebra/Order/Archimedean/Basic.lean +++ b/Mathlib/Algebra/Order/Archimedean/Basic.lean @@ -3,6 +3,7 @@ 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.Algebra.Order.Group.Units import Mathlib.Algebra.Order.Ring.Pow import Mathlib.Data.Int.LeastGreatest import Mathlib.Data.Rat.Floor diff --git a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean index e0e51858c4b0b..95813028bae02 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean @@ -7,9 +7,8 @@ import Mathlib.Algebra.GroupWithZero.InjSurj import Mathlib.Algebra.GroupWithZero.Units.Equiv import Mathlib.Algebra.GroupWithZero.WithZero import Mathlib.Algebra.Order.AddGroupWithTop -import Mathlib.Algebra.Order.Group.Units import Mathlib.Algebra.Order.GroupWithZero.Synonym -import Mathlib.Algebra.Order.GroupWithZero.Unbundled +import Mathlib.Algebra.Order.GroupWithZero.Unbundled.Lemmas import Mathlib.Algebra.Order.Monoid.Basic import Mathlib.Algebra.Order.Monoid.OrderDual import Mathlib.Algebra.Order.Monoid.TypeTags @@ -130,96 +129,63 @@ instance (priority := 100) LinearOrderedCommGroupWithZero.toMulPosStrictMono : MulPosStrictMono α where elim a b c hbc := by by_contra! h; exact hbc.not_le <| (mul_le_mul_right a.2).1 h -/-- Alias of `one_le_mul'` for unification. -/ -@[deprecated one_le_mul (since := "2024-08-21")] -theorem one_le_mul₀ (ha : 1 ≤ a) (hb : 1 ≤ b) : 1 ≤ a * b := - one_le_mul ha hb - -@[deprecated mul_le_mul_right (since := "2024-08-21")] -theorem le_of_le_mul_right (h : c ≠ 0) (hab : a * c ≤ b * c) : a ≤ b := - (mul_le_mul_right (zero_lt_iff.2 h)).1 hab - -@[deprecated le_mul_inv_iff₀ (since := "2024-08-21")] -theorem le_mul_inv_of_mul_le (h : c ≠ 0) (hab : a * c ≤ b) : a ≤ b * c⁻¹ := - (le_mul_inv_iff₀ (zero_lt_iff.2 h)).2 hab - -theorem mul_inv_le_of_le_mul (hab : a ≤ b * c) : a * c⁻¹ ≤ b := by - by_cases h : c = 0 - · simp [h] - · exact (mul_le_mul_right (zero_lt_iff.2 h)).1 (by simpa [h] using hab) +@[deprecated mul_inv_le_of_le_mul₀ (since := "2024-11-18")] +theorem mul_inv_le_of_le_mul (hab : a ≤ b * c) : a * c⁻¹ ≤ b := + mul_inv_le_of_le_mul₀ zero_le' zero_le' hab @[simp] theorem Units.zero_lt (u : αˣ) : (0 : α) < u := - zero_lt_iff.2 <| u.ne_zero + zero_lt_iff.2 u.ne_zero +@[deprecated mul_lt_mul_of_le_of_lt_of_nonneg_of_pos (since := "2024-11-18")] theorem mul_lt_mul_of_lt_of_le₀ (hab : a ≤ b) (hb : b ≠ 0) (hcd : c < d) : a * c < b * d := - have hd : d ≠ 0 := ne_zero_of_lt hcd - if ha : a = 0 then by - rw [ha, zero_mul, zero_lt_iff] - exact mul_ne_zero hb hd - else - if hc : c = 0 then by - rw [hc, mul_zero, zero_lt_iff] - exact mul_ne_zero hb hd - else - show Units.mk0 a ha * Units.mk0 c hc < Units.mk0 b hb * Units.mk0 d hd from - mul_lt_mul_of_le_of_lt hab hcd + mul_lt_mul_of_le_of_lt_of_nonneg_of_pos hab hcd zero_le' (zero_lt_iff.2 hb) +@[deprecated mul_lt_mul'' (since := "2024-11-18")] theorem mul_lt_mul₀ (hab : a < b) (hcd : c < d) : a * c < b * d := - mul_lt_mul_of_lt_of_le₀ hab.le (ne_zero_of_lt hab) hcd + mul_lt_mul'' hab hcd zero_le' zero_le' theorem mul_inv_lt_of_lt_mul₀ (h : a < b * c) : a * c⁻¹ < b := by contrapose! h - simpa only [inv_inv] using mul_inv_le_of_le_mul h + simpa only [inv_inv] using mul_inv_le_of_le_mul₀ zero_le' zero_le' h theorem inv_mul_lt_of_lt_mul₀ (h : a < b * c) : b⁻¹ * a < c := by rw [mul_comm] at * exact mul_inv_lt_of_lt_mul₀ h -@[deprecated mul_lt_mul_of_pos_right (since := "2024-08-21")] -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 lt_of_mul_lt_mul_of_le₀ (h : a * b < c * d) (hc : 0 < c) (hh : c ≤ a) : b < d := by have ha : a ≠ 0 := ne_of_gt (lt_of_lt_of_le hc hh) rw [← inv_le_inv₀ (zero_lt_iff.2 ha) hc] at hh - have := mul_lt_mul_of_lt_of_le₀ hh (inv_ne_zero (ne_of_gt hc)) h - simpa [inv_mul_cancel_left₀ ha, inv_mul_cancel_left₀ (ne_of_gt hc)] using this + simpa [inv_mul_cancel_left₀ ha, inv_mul_cancel_left₀ hc.ne'] + using mul_lt_mul_of_le_of_lt_of_nonneg_of_pos hh h zero_le' (inv_pos.2 hc) -@[deprecated mul_le_mul_right (since := "2024-08-21")] -theorem mul_le_mul_right₀ (hc : c ≠ 0) : a * c ≤ b * c ↔ a ≤ b := - mul_le_mul_right (zero_lt_iff.2 hc) +@[deprecated div_le_div_iff_of_pos_right (since := "2024-11-18")] +theorem div_le_div_right₀ (hc : c ≠ 0) : a / c ≤ b / c ↔ a ≤ b := + div_le_div_iff_of_pos_right (zero_lt_iff.2 hc) -@[deprecated mul_le_mul_left (since := "2024-08-21")] -theorem mul_le_mul_left₀ (ha : a ≠ 0) : a * b ≤ a * c ↔ b ≤ c := - mul_le_mul_left (zero_lt_iff.2 ha) +@[deprecated div_le_div_iff_of_pos_left (since := "2024-11-18")] +theorem div_le_div_left₀ (ha : a ≠ 0) (hb : b ≠ 0) (hc : c ≠ 0) : a / b ≤ a / c ↔ c ≤ b := + div_le_div_iff_of_pos_left (zero_lt_iff.2 ha) (zero_lt_iff.2 hb) (zero_lt_iff.2 hc) -theorem div_le_div_right₀ (hc : c ≠ 0) : a / c ≤ b / c ↔ a ≤ b := by - rw [div_eq_mul_inv, div_eq_mul_inv, mul_le_mul_right (zero_lt_iff.2 (inv_ne_zero hc))] - -theorem div_le_div_left₀ (ha : a ≠ 0) (hb : b ≠ 0) (hc : c ≠ 0) : a / b ≤ a / c ↔ c ≤ b := by - simp only [div_eq_mul_inv, mul_le_mul_left (zero_lt_iff.2 ha), - inv_le_inv₀ (zero_lt_iff.2 hb) (zero_lt_iff.2 hc)] - -/-- `Equiv.mulLeft₀` as an `OrderIso` on a `LinearOrderedCommGroupWithZero.`. - -Note that `OrderIso.mulLeft₀` refers to the `LinearOrderedField` version. -/ -@[simps! (config := { simpRhs := true }) apply toEquiv] -def OrderIso.mulLeft₀' {a : α} (ha : a ≠ 0) : α ≃o α := - { Equiv.mulLeft₀ a ha with map_rel_iff' := mul_le_mul_left (zero_lt_iff.2 ha) } +/-- `Equiv.mulLeft₀` as an `OrderIso` on a `LinearOrderedCommGroupWithZero.`. -/ +@[simps! (config := { simpRhs := true }) apply toEquiv, +deprecated OrderIso.mulLeft₀ (since := "2024-11-18")] +def OrderIso.mulLeft₀' {a : α} (ha : a ≠ 0) : α ≃o α := .mulLeft₀ a (zero_lt_iff.2 ha) +set_option linter.deprecated false in +@[deprecated OrderIso.mulLeft₀_symm (since := "2024-11-18")] theorem OrderIso.mulLeft₀'_symm {a : α} (ha : a ≠ 0) : (OrderIso.mulLeft₀' ha).symm = OrderIso.mulLeft₀' (inv_ne_zero ha) := by ext rfl -/-- `Equiv.mulRight₀` as an `OrderIso` on a `LinearOrderedCommGroupWithZero.`. - -Note that `OrderIso.mulRight₀` refers to the `LinearOrderedField` version. -/ -@[simps! (config := { simpRhs := true }) apply toEquiv] -def OrderIso.mulRight₀' {a : α} (ha : a ≠ 0) : α ≃o α := - { Equiv.mulRight₀ a ha with map_rel_iff' := mul_le_mul_right (zero_lt_iff.2 ha) } +/-- `Equiv.mulRight₀` as an `OrderIso` on a `LinearOrderedCommGroupWithZero.`. -/ +@[simps! (config := { simpRhs := true }) apply toEquiv, +deprecated OrderIso.mulRight₀ (since := "2024-11-18")] +def OrderIso.mulRight₀' {a : α} (ha : a ≠ 0) : α ≃o α := .mulRight₀ a (zero_lt_iff.2 ha) +set_option linter.deprecated false in +@[deprecated OrderIso.mulRight₀_symm (since := "2024-11-18")] theorem OrderIso.mulRight₀'_symm {a : α} (ha : a ≠ 0) : (OrderIso.mulRight₀' ha).symm = OrderIso.mulRight₀' (inv_ne_zero ha) := by ext @@ -231,9 +197,8 @@ instance : LinearOrderedAddCommGroupWithTop (Additive αᵒᵈ) := neg_top := inv_zero (G₀ := α) add_neg_cancel := fun a ha ↦ mul_inv_cancel₀ (G₀ := α) (id ha : Additive.toMul a ≠ 0) } -lemma pow_lt_pow_succ (ha : 1 < a) : a ^ n < a ^ n.succ := by - rw [← one_mul (a ^ n), pow_succ'] - exact mul_lt_mul_of_pos_right ha (pow_pos (zero_lt_one.trans ha) _) +@[deprecated pow_lt_pow_right₀ (since := "2024-11-18")] +lemma pow_lt_pow_succ (ha : 1 < a) : a ^ n < a ^ n.succ := pow_lt_pow_right₀ ha n.lt_succ_self end LinearOrderedCommGroupWithZero diff --git a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean index 431015c2b62e0..bbd7e7951179d 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean @@ -1724,7 +1724,6 @@ lemma div_lt_div₀ (hac : a < c) (hdb : d ≤ b) (hc : 0 ≤ c) (hd : 0 < d) : rw [div_eq_mul_inv, div_eq_mul_inv] exact mul_lt_mul hac ((inv_le_inv₀ (hd.trans_le hdb) hd).2 hdb) (inv_pos.2 <| hd.trans_le hdb) hc -/-- See -/ lemma div_lt_div₀' (hac : a ≤ c) (hdb : d < b) (hc : 0 < c) (hd : 0 < d) : a / b < c / d := by rw [div_eq_mul_inv, div_eq_mul_inv] exact mul_lt_mul' hac ((inv_lt_inv₀ (hd.trans hdb) hd).2 hdb) diff --git a/Mathlib/Algebra/Order/GroupWithZero/Unbundled/Lemmas.lean b/Mathlib/Algebra/Order/GroupWithZero/Unbundled/Lemmas.lean index 15331ca9c673a..daa1991832ae7 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Unbundled/Lemmas.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Unbundled/Lemmas.lean @@ -12,26 +12,48 @@ import Mathlib.Order.Hom.Basic # Multiplication by a positive element as an order isomorphism -/ -variable {G₀} [GroupWithZero G₀] [Preorder G₀] {a : G₀} +namespace OrderIso +variable {G₀} [GroupWithZero G₀] [PartialOrder G₀] + +section left +variable [PosMulMono G₀] [PosMulReflectLE 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 +def mulLeft₀ (a : G₀) (ha : 0 < a) : G₀ ≃o G₀ where toEquiv := .mulLeft₀ a ha.ne' map_rel_iff' := mul_le_mul_left ha +variable [ZeroLEOneClass G₀] + +lemma mulLeft₀_symm (a : G₀) (ha : 0 < a) : (mulLeft₀ a ha).symm = mulLeft₀ a⁻¹ (inv_pos.2 ha) := by + ext; rfl + +end left + +section right +variable [MulPosMono G₀] [MulPosReflectLE G₀] + /-- `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 +def mulRight₀ (a : G₀) (ha : 0 < a) : G₀ ≃o G₀ where toEquiv := .mulRight₀ a ha.ne' map_rel_iff' := mul_le_mul_right ha +variable [ZeroLEOneClass G₀] [PosMulReflectLT G₀] + +lemma mulRight₀_symm (a : G₀) (ha : 0 < a) : + (mulRight₀ a ha).symm = mulRight₀ a⁻¹ (inv_pos.2 ha) := by ext; rfl + +end right + /-- `Equiv.divRight₀` as an order isomorphism. -/ @[simps! (config := { simpRhs := true })] -def OrderIso.divRight₀ {G₀} [GroupWithZero G₀] [PartialOrder G₀] [ZeroLEOneClass G₀] - [MulPosStrictMono G₀] [MulPosReflectLE G₀] [PosMulReflectLT G₀] +def divRight₀ [ZeroLEOneClass G₀] [MulPosStrictMono G₀] [MulPosReflectLE G₀] [PosMulReflectLT G₀] (a : G₀) (ha : 0 < a) : G₀ ≃o G₀ where toEquiv := .divRight₀ a ha.ne' map_rel_iff' {b c} := by simp only [Equiv.divRight₀_apply, div_eq_mul_inv] exact mul_le_mul_right (a := a⁻¹) (inv_pos.mpr ha) + +end OrderIso diff --git a/Mathlib/Algebra/Order/Hom/Monoid.lean b/Mathlib/Algebra/Order/Hom/Monoid.lean index 5fa19423317a2..b40dacaccb13e 100644 --- a/Mathlib/Algebra/Order/Hom/Monoid.lean +++ b/Mathlib/Algebra/Order/Hom/Monoid.lean @@ -6,6 +6,7 @@ Authors: Yaël Dillies import Mathlib.Algebra.GroupWithZero.Hom import Mathlib.Algebra.Order.Group.Instances import Mathlib.Algebra.Order.GroupWithZero.Canonical +import Mathlib.Algebra.Order.Monoid.Units import Mathlib.Order.Hom.Basic /-! diff --git a/Mathlib/RingTheory/LaurentSeries.lean b/Mathlib/RingTheory/LaurentSeries.lean index 2d9c6f7c2f02e..d543b3fa95654 100644 --- a/Mathlib/RingTheory/LaurentSeries.lean +++ b/Mathlib/RingTheory/LaurentSeries.lean @@ -911,7 +911,8 @@ theorem exists_Polynomial_intValuation_lt (F : K⟦X⟧) (η : ℤₘ₀ˣ) : rw [← valuation_of_algebraMap (K := K⸨X⸩) (PowerSeries.idealX K) (F - F.trunc (d + 1))] apply lt_of_le_of_lt this rw [← mul_one (η : ℤₘ₀), mul_assoc, one_mul] - apply mul_lt_mul_of_lt_of_le₀ (le_refl _) η.ne_zero + gcongr + · exact zero_lt_iff.2 η.ne_zero rw [← WithZero.coe_one, coe_lt_coe, ofAdd_neg, Right.inv_lt_one_iff, ← ofAdd_zero, Multiplicative.ofAdd_lt] exact Int.zero_lt_one diff --git a/Mathlib/RingTheory/Perfection.lean b/Mathlib/RingTheory/Perfection.lean index 013131970238e..7db1bdda6d237 100644 --- a/Mathlib/RingTheory/Perfection.lean +++ b/Mathlib/RingTheory/Perfection.lean @@ -424,7 +424,7 @@ theorem mul_ne_zero_of_pow_p_ne_zero {x y : ModP K v O hv p} (hx : x ^ p ≠ 0) rw [← v_p_lt_val hv] at hx hy ⊢ rw [RingHom.map_pow, v.map_pow, ← rpow_lt_rpow_iff h1p, ← rpow_natCast, ← rpow_mul, mul_one_div_cancel (Nat.cast_ne_zero.2 hp.1.ne_zero : (p : ℝ) ≠ 0), rpow_one] at hx hy - rw [RingHom.map_mul, v.map_mul]; refine lt_of_le_of_lt ?_ (mul_lt_mul₀ hx hy) + rw [RingHom.map_mul, v.map_mul]; refine lt_of_le_of_lt ?_ (mul_lt_mul'' hx hy zero_le' zero_le') by_cases hvp : v p = 0 · rw [hvp]; exact zero_le _ replace hvp := zero_lt_iff.2 hvp diff --git a/Mathlib/RingTheory/Valuation/Integral.lean b/Mathlib/RingTheory/Valuation/Integral.lean index c8e1a11e3d970..515de9eca405b 100644 --- a/Mathlib/RingTheory/Valuation/Integral.lean +++ b/Mathlib/RingTheory/Valuation/Integral.lean @@ -37,7 +37,7 @@ theorem mem_of_integral {x : R} (hx : IsIntegral O x) : x ∈ v.integer := rw [eval₂_mul, eval₂_pow, eval₂_C, eval₂_X, v.map_mul, v.map_pow, ← one_mul (v x ^ p.natDegree)] cases' (hv.2 <| p.coeff i).lt_or_eq with hvpi hvpi - · exact mul_lt_mul₀ hvpi (pow_lt_pow_right₀ hvx <| Finset.mem_range.1 hi) + · exact mul_lt_mul'' hvpi (pow_lt_pow_right₀ hvx <| Finset.mem_range.1 hi) zero_le' zero_le' · rw [hvpi, one_mul, one_mul]; exact pow_lt_pow_right₀ hvx (Finset.mem_range.1 hi) protected theorem integralClosure : integralClosure O R = ⊥ := diff --git a/Mathlib/RingTheory/Valuation/ValuationSubring.lean b/Mathlib/RingTheory/Valuation/ValuationSubring.lean index a5ea654a7ca32..867c268d02ac5 100644 --- a/Mathlib/RingTheory/Valuation/ValuationSubring.lean +++ b/Mathlib/RingTheory/Valuation/ValuationSubring.lean @@ -480,7 +480,8 @@ section nonunits def nonunits : Subsemigroup K where carrier := {x | A.valuation x < 1} -- Porting note: added `Set.mem_setOf.mp` - mul_mem' ha hb := (mul_lt_mul₀ (Set.mem_setOf.mp ha) (Set.mem_setOf.mp hb)).trans_eq <| mul_one _ + mul_mem' ha hb := (mul_lt_mul'' (Set.mem_setOf.mp ha) (Set.mem_setOf.mp hb) + zero_le' zero_le').trans_eq <| mul_one _ theorem mem_nonunits_iff {x : K} : x ∈ A.nonunits ↔ A.valuation x < 1 := Iff.rfl diff --git a/Mathlib/Topology/Algebra/Valued/ValuationTopology.lean b/Mathlib/Topology/Algebra/Valued/ValuationTopology.lean index 8cdff99b6d7fe..621c4dd2159b2 100644 --- a/Mathlib/Topology/Algebra/Valued/ValuationTopology.lean +++ b/Mathlib/Topology/Algebra/Valued/ValuationTopology.lean @@ -42,9 +42,10 @@ theorem subgroups_basis : RingSubgroupsBasis fun γ : Γ₀ˣ => (v.ltAddSubgrou cases' exists_square_le γ with γ₀ h use γ₀ rintro - ⟨r, r_in, s, s_in, rfl⟩ + simp only [ltAddSubgroup, AddSubgroup.coe_set_mk, mem_setOf_eq] at r_in s_in calc (v (r * s) : Γ₀) = v r * v s := Valuation.map_mul _ _ _ - _ < γ₀ * γ₀ := mul_lt_mul₀ r_in s_in + _ < γ₀ * γ₀ := by gcongr <;> exact zero_le' _ ≤ γ := mod_cast h leftMul := by rintro x γ diff --git a/Mathlib/Topology/Algebra/WithZeroTopology.lean b/Mathlib/Topology/Algebra/WithZeroTopology.lean index c115bccac2248..cd47b1c9d0f8f 100644 --- a/Mathlib/Topology/Algebra/WithZeroTopology.lean +++ b/Mathlib/Topology/Algebra/WithZeroTopology.lean @@ -166,7 +166,7 @@ scoped instance (priority := 100) : ContinuousMul Γ₀ where refine ((hasBasis_nhds_zero.prod_nhds hasBasis_nhds_zero).tendsto_iff hasBasis_nhds_zero).2 fun γ hγ => ⟨(γ, 1), ⟨hγ, one_ne_zero⟩, ?_⟩ rintro ⟨x, y⟩ ⟨hx : x < γ, hy : y < 1⟩ - exact (mul_lt_mul₀ hx hy).trans_eq (mul_one γ) + exact (mul_lt_mul'' hx hy zero_le' zero_le').trans_eq (mul_one γ) · rw [zero_mul, nhds_prod_eq, nhds_of_ne_zero hy, prod_pure, tendsto_map'_iff] refine (hasBasis_nhds_zero.tendsto_iff hasBasis_nhds_zero).2 fun γ hγ => ?_ refine ⟨γ / y, div_ne_zero hγ hy, fun x hx => ?_⟩ From 7d4e3bb74bb8562fde3e5bf3f0157f0758e91832 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 20 Nov 2024 11:42:56 +0000 Subject: [PATCH 138/829] chore(discover-lean-pr-testing): properly fetch tags (#19204) Instead of guessing the dates of the latest two tags in the `lean4-nightly` repo, we list the available remote tags and take the last two. We then fetch them with a shallow depth, to avoid cloning the entire repo. After that we parse the commit log between these two tags. --- .../workflows/discover-lean-pr-testing.yml | 33 +++++++++---------- 1 file changed, 15 insertions(+), 18 deletions(-) diff --git a/.github/workflows/discover-lean-pr-testing.yml b/.github/workflows/discover-lean-pr-testing.yml index 9cc2434aea615..3c414d9e99689 100644 --- a/.github/workflows/discover-lean-pr-testing.yml +++ b/.github/workflows/discover-lean-pr-testing.yml @@ -24,19 +24,7 @@ jobs: git config --global user.name "github-actions" git config --global user.email "github-actions@github.com" - - name: Read date from lean-toolchain - id: fetch-date - run: | - TODAY=$(grep -oP 'nightly-\K\d{4}-\d{2}-\d{2}' lean-toolchain) - echo "TODAY=$TODAY" >> "$GITHUB_ENV" - - - name: Set YESTERDAY - id: set-yesterday - run: | - YESTERDAY=$(date -d "$TODAY -1 day" +%Y-%m-%d) - echo "YESTERDAY=$YESTERDAY" >> "$GITHUB_ENV" - - - name: Clone lean4-nightly repository and get PRs + - name: Clone lean4 repository and get PRs id: get-prs run: | NIGHTLY_URL="https://github.com/leanprover/lean4-nightly.git" @@ -44,14 +32,22 @@ jobs: # Create a temporary directory for cloning cd "$(mktemp -d)" || exit 1 - # Clone the repository with a depth of 30 (adjust as needed) - git clone --depth 30 "$NIGHTLY_URL" + # Clone the repository with a depth of 1 + git clone --depth 1 "$NIGHTLY_URL" # Navigate to the cloned repository cd lean4-nightly || exit 1 - YESTERDAY="${{ steps.set-yesterday.outputs.yesterday }}" - TODAY="${{ steps.fetch-date.outputs.today }}" - PRS=$(git log --oneline "nightly-$YESTERDAY..nightly-$TODAY" | sed 's/.*(#\([0-9]\+\))$/\1/') + + # Shallow fetch of the last two tags + LAST_TWO_TAGS=$(git ls-remote --tags | grep nightly | tail -2 | cut -f2) + OLD=$(echo "$LAST_TWO_TAGS" | head -1) + NEW=$(echo "$LAST_TWO_TAGS" | tail -1) + git fetch --depth=1 origin tag "$OLD" --no-tags + # Fetch the latest 100 commits prior to the $NEW tag. + # This should cover the terrain between $OLD and $NEW while still being a speedy download. + git fetch --depth=100 origin tag "$NEW" --no-tags + + PRS=$(git log --oneline "$OLD..$NEW" | sed 's/.*(#\([0-9]\+\))$/\1/') # Output the PRs echo "$PRS" @@ -64,6 +60,7 @@ jobs: PRS="${{ steps.get-prs.outputs.prs }}" echo "$PRS" | tr ' ' '\n' > prs.txt MATCHING_BRANCHES=$(git branch -r | grep -f prs.txt) + echo "$MATCHING_BRANCHES" # Initialize an empty variable to store branches with relevant diffs RELEVANT_BRANCHES="" From 7600042bca03f6a9066a2607026c72cfc6c034a5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 20 Nov 2024 11:42:58 +0000 Subject: [PATCH 139/829] feat: `IsScalarTower` and `SMulCommClass` instances for pointwise actions of submodules (#19214) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Moves: - `Submodule.smul_def` → `Submodule.setSemiring_smul_def` From GrowthInGroups Co-authored-by: Eric Wieser --- Mathlib/Algebra/Algebra/Operations.lean | 28 ++++++++++++++++--- .../Algebra/Module/Submodule/Pointwise.lean | 2 ++ 2 files changed, 26 insertions(+), 4 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Operations.lean b/Mathlib/Algebra/Algebra/Operations.lean index 8decae0eb7f39..de1a4621d5777 100644 --- a/Mathlib/Algebra/Algebra/Operations.lean +++ b/Mathlib/Algebra/Algebra/Operations.lean @@ -285,12 +285,14 @@ theorem mul_eq_map₂ : M * N = map₂ (LinearMap.mul R A) M N := le_antisymm (mul_le.mpr fun _m hm _n ↦ apply_mem_map₂ _ hm) (map₂_le.mpr fun _m hm _n ↦ mul_mem_mul hm) -variable (R) +variable (R M N) theorem span_mul_span : span R S * span R T = span R (S * T) := by rw [mul_eq_map₂]; apply map₂_span_span -variable {R} (M N P Q) +lemma mul_def : M * N = span R (M * N : Set A) := by simp [← span_mul_span] + +variable {R} (P Q) protected theorem mul_one : M * 1 = M := by conv_lhs => rw [one_eq_span, ← span_eq M] @@ -353,6 +355,24 @@ theorem comap_op_mul (M N : Submodule R Aᵐᵒᵖ) : comap (↑(opLinearEquiv R : A ≃ₗ[R] Aᵐᵒᵖ) : A →ₗ[R] Aᵐᵒᵖ) M := by simp_rw [comap_equiv_eq_map_symm, map_unop_mul] +section +variable {α : Type*} [Monoid α] [DistribMulAction α A] [SMulCommClass α R A] + +instance [IsScalarTower α A A] : IsScalarTower α (Submodule R A) (Submodule R A) where + smul_assoc a S T := by + rw [← S.span_eq, ← T.span_eq, smul_span, smul_eq_mul, smul_eq_mul, span_mul_span, span_mul_span, + smul_span, smul_mul_assoc] + +instance [SMulCommClass α A A] : SMulCommClass α (Submodule R A) (Submodule R A) where + smul_comm a S T := by + rw [← S.span_eq, ← T.span_eq, smul_span, smul_eq_mul, smul_eq_mul, span_mul_span, span_mul_span, + smul_span, mul_smul_comm] + +instance [SMulCommClass A α A] : SMulCommClass (Submodule R A) α (Submodule R A) := + have := SMulCommClass.symm A α A; .symm .. + +end + section open Pointwise @@ -635,7 +655,7 @@ instance moduleSet : Module (SetSemiring A) (Submodule R A) where variable {R A} -theorem smul_def (s : SetSemiring A) (P : Submodule R A) : +theorem setSemiring_smul_def (s : SetSemiring A) (P : Submodule R A) : s • P = span R (SetSemiring.down (α := A) s) * P := rfl @@ -647,7 +667,7 @@ theorem smul_le_smul {s t : SetSemiring A} {M N : Submodule R A} theorem singleton_smul (a : A) (M : Submodule R A) : Set.up ({a} : Set A) • M = M.map (LinearMap.mulLeft R a) := by conv_lhs => rw [← span_eq M] - rw [smul_def, SetSemiring.down_up, span_mul_span, singleton_mul] + rw [setSemiring_smul_def, SetSemiring.down_up, span_mul_span, singleton_mul] exact (map (LinearMap.mulLeft R a) M).span_eq section Quotient diff --git a/Mathlib/Algebra/Module/Submodule/Pointwise.lean b/Mathlib/Algebra/Module/Submodule/Pointwise.lean index 3e0acab091ae2..ec25efdc44f71 100644 --- a/Mathlib/Algebra/Module/Submodule/Pointwise.lean +++ b/Mathlib/Algebra/Module/Submodule/Pointwise.lean @@ -228,6 +228,8 @@ theorem smul_sup' (a : α) (S T : Submodule R M) : a • (S ⊔ T) = a • S ⊔ theorem smul_span (a : α) (s : Set M) : a • span R s = span R (a • s) := map_span _ _ +lemma smul_def (a : α) (S : Submodule R M) : a • S = span R (a • S : Set M) := by simp [← smul_span] + theorem span_smul (a : α) (s : Set M) : span R (a • s) = a • span R s := Eq.symm (span_image _).symm From 333a8fa711db4e6b0add704bc4b4bd55656a3f44 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 20 Nov 2024 11:42:59 +0000 Subject: [PATCH 140/829] feat: `Nat.cast_nonpos` (#19247) For simp confluence From GrowthInGroups --- Mathlib/Data/Int/Defs.lean | 3 ++- Mathlib/Data/Nat/Cast/Order/Basic.lean | 2 ++ 2 files changed, 4 insertions(+), 1 deletion(-) diff --git a/Mathlib/Data/Int/Defs.lean b/Mathlib/Data/Int/Defs.lean index 6bf06c9ff023a..9fd2c581ce8e9 100644 --- a/Mathlib/Data/Int/Defs.lean +++ b/Mathlib/Data/Int/Defs.lean @@ -114,7 +114,8 @@ lemma natCast_ne_zero_iff_pos {n : ℕ} : (n : ℤ) ≠ 0 ↔ 0 < n := by omega lemma natCast_succ_pos (n : ℕ) : 0 < (n.succ : ℤ) := natCast_pos.2 n.succ_pos -@[simp] lemma natCast_nonpos_iff {n : ℕ} : (n : ℤ) ≤ 0 ↔ n = 0 := by omega +-- We want to use this lemma earlier than the lemmas simp can prove it with +@[simp, nolint simpNF] lemma natCast_nonpos_iff {n : ℕ} : (n : ℤ) ≤ 0 ↔ n = 0 := by omega lemma natCast_nonneg (n : ℕ) : 0 ≤ (n : ℤ) := ofNat_le.2 (Nat.zero_le _) diff --git a/Mathlib/Data/Nat/Cast/Order/Basic.lean b/Mathlib/Data/Nat/Cast/Order/Basic.lean index 21129456cd441..eec03f2914acb 100644 --- a/Mathlib/Data/Nat/Cast/Order/Basic.lean +++ b/Mathlib/Data/Nat/Cast/Order/Basic.lean @@ -98,6 +98,8 @@ theorem cast_lt_one : (n : α) < 1 ↔ n = 0 := by @[simp, norm_cast] theorem cast_le_one : (n : α) ≤ 1 ↔ n ≤ 1 := by rw [← cast_one, cast_le] +@[simp] lemma cast_nonpos : (n : α) ≤ 0 ↔ n = 0 := by norm_cast; omega + section variable [m.AtLeastTwo] From b6fb044ac0aeee41e114dea81dfa74126bde3aba Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Wed, 20 Nov 2024 11:43:00 +0000 Subject: [PATCH 141/829] chore(RingTheory): split `UniqueFactorizationDomain.lean` (#19256) This PR splits the big file `UniqueFactorizationDomain.lean` into many smaller files: * `UniqueFactorizationDomain.Defs` contains the definition of `WfDvdMonoid`, `UniqueFactorizationMonoid` and `factors` * `UniqueFactorizationDomain.Basic` contains basic results on the previous definition such as the equivalence between `UniqueFactorizationMonoid` and the existence of prime factors * `UniqueFactorizationDomain.NormalizedFactors` defines `UniqueFactorizationMonoid.normalizedFactors` and contains basic results. (Or should we merge it into `Defs`/`Basic`?) * `UniqueFactorizationDomain.FactorSet` defines `Associates.FactorSet` and `Associates.factors` * `UniqueFactorizationDomain.Finite` proves that elements of a UFM with finitely many units have finitely many divisors. (This could be moved to `Basic.lean` but I find it niche enough to deserve its own file.) * `UniqueFactorizationDomain.Finsupp` defines `UniqueFactorizationMonoid.factorization` as a `Finsupp` version of `factors` * `UniqueFactorizationDomain.GCDMonoid` defines a GCD on a UFM * `UniqueFactorizationDomain.Ideal` shows UFDs satisfy the ascending chain condition on ideals * `UniqueFactorizationDomain.Multiplicative` contains results on multiplicative properties and functions. * `UniqueFactorizationDomain.Multiplicity` relates `factors.count` and `multiplicity` * `UniqueFactorizationDomain.Nat` contains the `UniqueFactorizationMonoid Nat` instance --- Mathlib.lean | 12 +- Mathlib/Algebra/Squarefree/Basic.lean | 3 +- .../Additive/ErdosGinzburgZiv.lean | 1 - Mathlib/Data/Nat/Squarefree.lean | 1 + Mathlib/NumberTheory/ArithmeticFunction.lean | 4 +- Mathlib/NumberTheory/FLT/Four.lean | 1 + Mathlib/RingTheory/ChainOfDivisors.lean | 3 +- .../DedekindDomain/Factorization.lean | 1 + .../DiscreteValuationRing/Basic.lean | 1 + Mathlib/RingTheory/Ideal/Norm/AbsNorm.lean | 1 + .../RingTheory/Localization/Away/Basic.lean | 2 +- Mathlib/RingTheory/Localization/NumDen.lean | 2 +- .../Noetherian/UniqueFactorizationDomain.lean | 2 +- Mathlib/RingTheory/Polynomial/Radical.lean | 1 + .../Polynomial/UniqueFactorization.lean | 3 + Mathlib/RingTheory/PowerSeries/Inverse.lean | 1 + Mathlib/RingTheory/PrincipalIdealDomain.lean | 1 + Mathlib/RingTheory/Radical.lean | 3 +- Mathlib/RingTheory/RootsOfUnity/Minpoly.lean | 5 +- .../RingTheory/UniqueFactorizationDomain.lean | 2017 ----------------- .../UniqueFactorizationDomain/Basic.lean | 473 ++++ .../UniqueFactorizationDomain/Defs.lean | 205 ++ .../UniqueFactorizationDomain/FactorSet.lean | 647 ++++++ .../UniqueFactorizationDomain/Finite.lean | 53 + .../UniqueFactorizationDomain/Finsupp.lean | 62 + .../UniqueFactorizationDomain/GCDMonoid.lean | 73 + .../UniqueFactorizationDomain/Ideal.lean | 58 + .../Multiplicative.lean | 163 ++ .../Multiplicity.lean | 126 + .../UniqueFactorizationDomain/Nat.lean | 62 + .../NormalizedFactors.lean | 329 +++ MathlibTest/Variable.lean | 8 +- 32 files changed, 2292 insertions(+), 2032 deletions(-) delete mode 100644 Mathlib/RingTheory/UniqueFactorizationDomain.lean create mode 100644 Mathlib/RingTheory/UniqueFactorizationDomain/Basic.lean create mode 100644 Mathlib/RingTheory/UniqueFactorizationDomain/Defs.lean create mode 100644 Mathlib/RingTheory/UniqueFactorizationDomain/FactorSet.lean create mode 100644 Mathlib/RingTheory/UniqueFactorizationDomain/Finite.lean create mode 100644 Mathlib/RingTheory/UniqueFactorizationDomain/Finsupp.lean create mode 100644 Mathlib/RingTheory/UniqueFactorizationDomain/GCDMonoid.lean create mode 100644 Mathlib/RingTheory/UniqueFactorizationDomain/Ideal.lean create mode 100644 Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicative.lean create mode 100644 Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean create mode 100644 Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean create mode 100644 Mathlib/RingTheory/UniqueFactorizationDomain/NormalizedFactors.lean diff --git a/Mathlib.lean b/Mathlib.lean index f0d237abd75a7..0426236be0a65 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4467,7 +4467,17 @@ import Mathlib.RingTheory.TwoSidedIdeal.Instances import Mathlib.RingTheory.TwoSidedIdeal.Kernel import Mathlib.RingTheory.TwoSidedIdeal.Lattice import Mathlib.RingTheory.TwoSidedIdeal.Operations -import Mathlib.RingTheory.UniqueFactorizationDomain +import Mathlib.RingTheory.UniqueFactorizationDomain.Basic +import Mathlib.RingTheory.UniqueFactorizationDomain.Defs +import Mathlib.RingTheory.UniqueFactorizationDomain.FactorSet +import Mathlib.RingTheory.UniqueFactorizationDomain.Finite +import Mathlib.RingTheory.UniqueFactorizationDomain.Finsupp +import Mathlib.RingTheory.UniqueFactorizationDomain.GCDMonoid +import Mathlib.RingTheory.UniqueFactorizationDomain.Ideal +import Mathlib.RingTheory.UniqueFactorizationDomain.Multiplicative +import Mathlib.RingTheory.UniqueFactorizationDomain.Multiplicity +import Mathlib.RingTheory.UniqueFactorizationDomain.Nat +import Mathlib.RingTheory.UniqueFactorizationDomain.NormalizedFactors import Mathlib.RingTheory.Unramified.Basic import Mathlib.RingTheory.Unramified.Field import Mathlib.RingTheory.Unramified.Finite diff --git a/Mathlib/Algebra/Squarefree/Basic.lean b/Mathlib/Algebra/Squarefree/Basic.lean index 284bc77d8edec..9d68d11fc173a 100644 --- a/Mathlib/Algebra/Squarefree/Basic.lean +++ b/Mathlib/Algebra/Squarefree/Basic.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ import Mathlib.RingTheory.Nilpotent.Basic -import Mathlib.RingTheory.UniqueFactorizationDomain +import Mathlib.RingTheory.UniqueFactorizationDomain.GCDMonoid +import Mathlib.RingTheory.UniqueFactorizationDomain.Multiplicity /-! # Squarefree elements of monoids diff --git a/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean b/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean index 0dc3527d3bad6..5d0d0f094055a 100644 --- a/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean +++ b/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean @@ -6,7 +6,6 @@ Authors: Yaël Dillies import Mathlib.Algebra.BigOperators.Ring import Mathlib.Data.Multiset.Fintype import Mathlib.FieldTheory.ChevalleyWarning -import Mathlib.RingTheory.UniqueFactorizationDomain /-! # The Erdős–Ginzburg–Ziv theorem diff --git a/Mathlib/Data/Nat/Squarefree.lean b/Mathlib/Data/Nat/Squarefree.lean index 40e3df63b05fc..53bb8bb29fec5 100644 --- a/Mathlib/Data/Nat/Squarefree.lean +++ b/Mathlib/Data/Nat/Squarefree.lean @@ -5,6 +5,7 @@ Authors: Aaron Anderson -/ import Mathlib.Algebra.Squarefree.Basic import Mathlib.Data.Nat.Factorization.PrimePow +import Mathlib.RingTheory.UniqueFactorizationDomain.Nat /-! # Lemmas about squarefreeness of natural numbers diff --git a/Mathlib/NumberTheory/ArithmeticFunction.lean b/Mathlib/NumberTheory/ArithmeticFunction.lean index d6a37afa99adb..5c13f9d53ebe2 100644 --- a/Mathlib/NumberTheory/ArithmeticFunction.lean +++ b/Mathlib/NumberTheory/ArithmeticFunction.lean @@ -299,7 +299,7 @@ theorem one_smul' (b : ArithmeticFunction M) : (1 : ArithmeticFunction R) • b intro y ymem ynmem have y1ne : y.fst ≠ 1 := by intro con - simp only [Con, mem_divisorsAntidiagonal, one_mul, Ne] at ymem + simp only [mem_divisorsAntidiagonal, one_mul, Ne] at ymem simp only [mem_singleton, Prod.ext_iff] at ynmem -- Porting note: `tauto` worked from here. cases y @@ -331,7 +331,7 @@ instance instMonoid : Monoid (ArithmeticFunction R) := have y2ne : y.snd ≠ 1 := by intro con cases y; subst con -- Porting note: added - simp only [Con, mem_divisorsAntidiagonal, mul_one, Ne] at ymem + simp only [mem_divisorsAntidiagonal, mul_one, Ne] at ymem simp only [mem_singleton, Prod.ext_iff] at ynmem tauto simp [y2ne] diff --git a/Mathlib/NumberTheory/FLT/Four.lean b/Mathlib/NumberTheory/FLT/Four.lean index e6da8e8d4814f..b6802f986775d 100644 --- a/Mathlib/NumberTheory/FLT/Four.lean +++ b/Mathlib/NumberTheory/FLT/Four.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Paul van Wamelen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Paul van Wamelen -/ +import Mathlib.Data.Nat.Factors import Mathlib.NumberTheory.FLT.Basic import Mathlib.NumberTheory.PythagoreanTriples import Mathlib.RingTheory.Coprime.Lemmas diff --git a/Mathlib/RingTheory/ChainOfDivisors.lean b/Mathlib/RingTheory/ChainOfDivisors.lean index f22b75532ab93..a1921d9752ce6 100644 --- a/Mathlib/RingTheory/ChainOfDivisors.lean +++ b/Mathlib/RingTheory/ChainOfDivisors.lean @@ -3,10 +3,11 @@ Copyright (c) 2021 Paul Lezeau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen, Paul Lezeau -/ +import Mathlib.Algebra.GCDMonoid.Basic import Mathlib.Algebra.IsPrimePow import Mathlib.Algebra.Squarefree.Basic -import Mathlib.Algebra.GCDMonoid.Basic import Mathlib.Data.ZMod.Defs +import Mathlib.Order.Atoms import Mathlib.Order.Hom.Bounded /-! diff --git a/Mathlib/RingTheory/DedekindDomain/Factorization.lean b/Mathlib/RingTheory/DedekindDomain/Factorization.lean index 9a837256ab9a2..61803b29f13c4 100644 --- a/Mathlib/RingTheory/DedekindDomain/Factorization.lean +++ b/Mathlib/RingTheory/DedekindDomain/Factorization.lean @@ -5,6 +5,7 @@ Authors: María Inés de Frutos-Fernández -/ import Mathlib.Order.Filter.Cofinite import Mathlib.RingTheory.DedekindDomain.Ideal +import Mathlib.RingTheory.UniqueFactorizationDomain.Finite /-! # Factorization of ideals and fractional ideals of Dedekind domains diff --git a/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean b/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean index cbcfd6561f1b7..33a536b988ff1 100644 --- a/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean +++ b/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean @@ -5,6 +5,7 @@ Authors: Kevin Buzzard -/ import Mathlib.RingTheory.AdicCompletion.Basic import Mathlib.RingTheory.LocalRing.MaximalIdeal.Basic +import Mathlib.RingTheory.UniqueFactorizationDomain.Basic import Mathlib.RingTheory.Valuation.PrimeMultiplicity import Mathlib.RingTheory.Valuation.ValuationRing diff --git a/Mathlib/RingTheory/Ideal/Norm/AbsNorm.lean b/Mathlib/RingTheory/Ideal/Norm/AbsNorm.lean index 69f9fc150ca4f..bf904f8bef341 100644 --- a/Mathlib/RingTheory/Ideal/Norm/AbsNorm.lean +++ b/Mathlib/RingTheory/Ideal/Norm/AbsNorm.lean @@ -11,6 +11,7 @@ import Mathlib.LinearAlgebra.FreeModule.IdealQuotient import Mathlib.RingTheory.DedekindDomain.Dvr import Mathlib.RingTheory.DedekindDomain.Ideal import Mathlib.RingTheory.Norm.Basic +import Mathlib.RingTheory.UniqueFactorizationDomain.Multiplicative /-! diff --git a/Mathlib/RingTheory/Localization/Away/Basic.lean b/Mathlib/RingTheory/Localization/Away/Basic.lean index 3f7cd7a8e49b6..d587619c72796 100644 --- a/Mathlib/RingTheory/Localization/Away/Basic.lean +++ b/Mathlib/RingTheory/Localization/Away/Basic.lean @@ -6,7 +6,7 @@ Authors: Kenny Lau, Mario Carneiro, Johan Commelin, Amelia Livingston, Anne Baan import Mathlib.GroupTheory.MonoidLocalization.Away import Mathlib.RingTheory.Ideal.Maps import Mathlib.RingTheory.Localization.Basic -import Mathlib.RingTheory.UniqueFactorizationDomain +import Mathlib.RingTheory.UniqueFactorizationDomain.Multiplicity /-! # Localizations away from an element diff --git a/Mathlib/RingTheory/Localization/NumDen.lean b/Mathlib/RingTheory/Localization/NumDen.lean index 9d72b32932ecc..290ca0e000441 100644 --- a/Mathlib/RingTheory/Localization/NumDen.lean +++ b/Mathlib/RingTheory/Localization/NumDen.lean @@ -5,7 +5,7 @@ Authors: Kenny Lau, Mario Carneiro, Johan Commelin, Amelia Livingston, Anne Baan -/ import Mathlib.RingTheory.Localization.FractionRing import Mathlib.RingTheory.Localization.Integer -import Mathlib.RingTheory.UniqueFactorizationDomain +import Mathlib.RingTheory.UniqueFactorizationDomain.GCDMonoid /-! # Numerator and denominator in a localization diff --git a/Mathlib/RingTheory/Noetherian/UniqueFactorizationDomain.lean b/Mathlib/RingTheory/Noetherian/UniqueFactorizationDomain.lean index a56bfca5ca1bf..b475161288512 100644 --- a/Mathlib/RingTheory/Noetherian/UniqueFactorizationDomain.lean +++ b/Mathlib/RingTheory/Noetherian/UniqueFactorizationDomain.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ import Mathlib.RingTheory.Noetherian.Defs -import Mathlib.RingTheory.UniqueFactorizationDomain +import Mathlib.RingTheory.UniqueFactorizationDomain.Ideal /-! # Noetherian domains have unique factorization diff --git a/Mathlib/RingTheory/Polynomial/Radical.lean b/Mathlib/RingTheory/Polynomial/Radical.lean index 71a22fad74761..4a8f0a57af1e6 100644 --- a/Mathlib/RingTheory/Polynomial/Radical.lean +++ b/Mathlib/RingTheory/Polynomial/Radical.lean @@ -6,6 +6,7 @@ Authors: Jineon Baek, Seewoo Lee import Mathlib.Algebra.Polynomial.FieldDivision import Mathlib.RingTheory.Polynomial.Wronskian import Mathlib.RingTheory.Radical +import Mathlib.RingTheory.UniqueFactorizationDomain.Multiplicative /-! # Radical of a polynomial diff --git a/Mathlib/RingTheory/Polynomial/UniqueFactorization.lean b/Mathlib/RingTheory/Polynomial/UniqueFactorization.lean index bf2b2308eda3b..6e24a23d71aee 100644 --- a/Mathlib/RingTheory/Polynomial/UniqueFactorization.lean +++ b/Mathlib/RingTheory/Polynomial/UniqueFactorization.lean @@ -5,6 +5,9 @@ Authors: Kenny Lau -/ import Mathlib.RingTheory.Polynomial.Basic import Mathlib.RingTheory.Polynomial.Content +import Mathlib.RingTheory.UniqueFactorizationDomain.Basic +import Mathlib.RingTheory.UniqueFactorizationDomain.Finite +import Mathlib.RingTheory.UniqueFactorizationDomain.GCDMonoid /-! # Unique factorization for univariate and multivariate polynomials diff --git a/Mathlib/RingTheory/PowerSeries/Inverse.lean b/Mathlib/RingTheory/PowerSeries/Inverse.lean index c8bef18f0aeb4..f69014616a2de 100644 --- a/Mathlib/RingTheory/PowerSeries/Inverse.lean +++ b/Mathlib/RingTheory/PowerSeries/Inverse.lean @@ -10,6 +10,7 @@ import Mathlib.RingTheory.MvPowerSeries.Inverse import Mathlib.RingTheory.PowerSeries.Basic import Mathlib.RingTheory.PowerSeries.Order import Mathlib.RingTheory.LocalRing.ResidueField.Defs +import Mathlib.RingTheory.UniqueFactorizationDomain.Multiplicity /-! # Formal power series - Inverses diff --git a/Mathlib/RingTheory/PrincipalIdealDomain.lean b/Mathlib/RingTheory/PrincipalIdealDomain.lean index 53c6974d87d3a..7cde766527865 100644 --- a/Mathlib/RingTheory/PrincipalIdealDomain.lean +++ b/Mathlib/RingTheory/PrincipalIdealDomain.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Morenikeji Neri -/ import Mathlib.Algebra.EuclideanDomain.Field +import Mathlib.Algebra.GCDMonoid.Basic import Mathlib.RingTheory.Ideal.Nonunits import Mathlib.RingTheory.Noetherian.UniqueFactorizationDomain diff --git a/Mathlib/RingTheory/Radical.lean b/Mathlib/RingTheory/Radical.lean index dbf09f0b12caa..a4558ba7c2ec9 100644 --- a/Mathlib/RingTheory/Radical.lean +++ b/Mathlib/RingTheory/Radical.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jineon Baek, Seewoo Lee -/ import Mathlib.Algebra.EuclideanDomain.Basic -import Mathlib.RingTheory.UniqueFactorizationDomain +import Mathlib.RingTheory.Coprime.Basic +import Mathlib.RingTheory.UniqueFactorizationDomain.NormalizedFactors /-! # Radical of an element of a unique factorization normalization monoid diff --git a/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean b/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean index 992c920e1efc9..34a8597d9b708 100644 --- a/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean +++ b/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean @@ -3,10 +3,11 @@ Copyright (c) 2020 Riccardo Brasca. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Riccardo Brasca, Johan Commelin -/ -import Mathlib.RingTheory.RootsOfUnity.PrimitiveRoots -import Mathlib.FieldTheory.Minpoly.IsIntegrallyClosed import Mathlib.Algebra.GCDMonoid.IntegrallyClosed import Mathlib.FieldTheory.Finite.Basic +import Mathlib.FieldTheory.Minpoly.IsIntegrallyClosed +import Mathlib.RingTheory.RootsOfUnity.PrimitiveRoots +import Mathlib.RingTheory.UniqueFactorizationDomain.Nat /-! # Minimal polynomial of roots of unity diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain.lean b/Mathlib/RingTheory/UniqueFactorizationDomain.lean deleted file mode 100644 index a5f452067a120..0000000000000 --- a/Mathlib/RingTheory/UniqueFactorizationDomain.lean +++ /dev/null @@ -1,2017 +0,0 @@ -/- -Copyright (c) 2018 Johannes Hölzl. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Johannes Hölzl, Jens Wagemaker, Aaron Anderson --/ -import Mathlib.Algebra.BigOperators.Associated -import Mathlib.Algebra.GCDMonoid.Basic -import Mathlib.Data.ENat.Lattice -import Mathlib.Data.Finsupp.Multiset -import Mathlib.Data.Nat.Factors -import Mathlib.RingTheory.Multiplicity -import Mathlib.RingTheory.Ideal.Operations - -/-! - -# Unique factorization - -## Main Definitions -* `WfDvdMonoid` holds for `Monoid`s for which a strict divisibility relation is - well-founded. -* `UniqueFactorizationMonoid` holds for `WfDvdMonoid`s where - `Irreducible` is equivalent to `Prime` - -## Main results -* `Ideal.setOf_isPrincipal_wellFoundedOn_gt`, `WfDvdMonoid.of_setOf_isPrincipal_wellFoundedOn_gt` - in a domain, well-foundedness of the strict version of ∣ is equivalent to the ascending - chain condition on principal ideals. - -## TODO -* set up the complete lattice structure on `FactorSet`. - --/ - - -variable {α : Type*} - -local infixl:50 " ~ᵤ " => Associated - -/-- Well-foundedness of the strict version of ∣, which is equivalent to the descending chain -condition on divisibility and to the ascending chain condition on -principal ideals in an integral domain. - -/ -abbrev WfDvdMonoid (α : Type*) [CommMonoidWithZero α] : Prop := - IsWellFounded α DvdNotUnit - -theorem wellFounded_dvdNotUnit {α : Type*} [CommMonoidWithZero α] [h : WfDvdMonoid α] : - WellFounded (DvdNotUnit (α := α)) := - h.wf - -namespace WfDvdMonoid - -variable [CommMonoidWithZero α] - -open Associates Nat - -theorem of_wfDvdMonoid_associates (_ : WfDvdMonoid (Associates α)) : WfDvdMonoid α := - ⟨(mk_surjective.wellFounded_iff mk_dvdNotUnit_mk_iff.symm).2 wellFounded_dvdNotUnit⟩ - -variable [WfDvdMonoid α] - -instance wfDvdMonoid_associates : WfDvdMonoid (Associates α) := - ⟨(mk_surjective.wellFounded_iff mk_dvdNotUnit_mk_iff.symm).1 wellFounded_dvdNotUnit⟩ - -theorem wellFoundedLT_associates : WellFoundedLT (Associates α) := - ⟨Subrelation.wf dvdNotUnit_of_lt wellFounded_dvdNotUnit⟩ - -@[deprecated wellFoundedLT_associates (since := "2024-09-02")] -theorem wellFounded_associates : WellFounded ((· < ·) : Associates α → Associates α → Prop) := - Subrelation.wf dvdNotUnit_of_lt wellFounded_dvdNotUnit - --- Porting note: elab_as_elim can only be global and cannot be changed on an imported decl --- attribute [local elab_as_elim] WellFounded.fix - -theorem exists_irreducible_factor {a : α} (ha : ¬IsUnit a) (ha0 : a ≠ 0) : - ∃ i, Irreducible i ∧ i ∣ a := - let ⟨b, hs, hr⟩ := wellFounded_dvdNotUnit.has_min { b | b ∣ a ∧ ¬IsUnit b } ⟨a, dvd_rfl, ha⟩ - ⟨b, - ⟨hs.2, fun c d he => - let h := dvd_trans ⟨d, he⟩ hs.1 - or_iff_not_imp_left.2 fun hc => - of_not_not fun hd => hr c ⟨h, hc⟩ ⟨ne_zero_of_dvd_ne_zero ha0 h, d, hd, he⟩⟩, - hs.1⟩ - -@[elab_as_elim] -theorem induction_on_irreducible {P : α → Prop} (a : α) (h0 : P 0) (hu : ∀ u : α, IsUnit u → P u) - (hi : ∀ a i : α, a ≠ 0 → Irreducible i → P a → P (i * a)) : P a := - haveI := Classical.dec - wellFounded_dvdNotUnit.fix - (fun a ih => - if ha0 : a = 0 then ha0.substr h0 - else - if hau : IsUnit a then hu a hau - else - let ⟨i, hii, b, hb⟩ := exists_irreducible_factor hau ha0 - let hb0 : b ≠ 0 := ne_zero_of_dvd_ne_zero ha0 ⟨i, mul_comm i b ▸ hb⟩ - hb.symm ▸ hi b i hb0 hii <| ih b ⟨hb0, i, hii.1, mul_comm i b ▸ hb⟩) - a - -theorem exists_factors (a : α) : - a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Irreducible b) ∧ Associated f.prod a := - induction_on_irreducible a (fun h => (h rfl).elim) - (fun _ hu _ => ⟨0, fun _ h => False.elim (Multiset.not_mem_zero _ h), hu.unit, one_mul _⟩) - fun a i ha0 hi ih _ => - let ⟨s, hs⟩ := ih ha0 - ⟨i ::ₘ s, fun b H => (Multiset.mem_cons.1 H).elim (fun h => h.symm ▸ hi) (hs.1 b), by - rw [s.prod_cons i] - exact hs.2.mul_left i⟩ - -theorem not_unit_iff_exists_factors_eq (a : α) (hn0 : a ≠ 0) : - ¬IsUnit a ↔ ∃ f : Multiset α, (∀ b ∈ f, Irreducible b) ∧ f.prod = a ∧ f ≠ ∅ := - ⟨fun hnu => by - obtain ⟨f, hi, u, rfl⟩ := exists_factors a hn0 - obtain ⟨b, h⟩ := Multiset.exists_mem_of_ne_zero fun h : f = 0 => hnu <| by simp [h] - classical - refine ⟨(f.erase b).cons (b * u), fun a ha => ?_, ?_, Multiset.cons_ne_zero⟩ - · obtain rfl | ha := Multiset.mem_cons.1 ha - exacts [Associated.irreducible ⟨u, rfl⟩ (hi b h), hi a (Multiset.mem_of_mem_erase ha)] - · rw [Multiset.prod_cons, mul_comm b, mul_assoc, Multiset.prod_erase h, mul_comm], - fun ⟨_, hi, he, hne⟩ => - let ⟨b, h⟩ := Multiset.exists_mem_of_ne_zero hne - not_isUnit_of_not_isUnit_dvd (hi b h).not_unit <| he ▸ Multiset.dvd_prod h⟩ - -theorem isRelPrime_of_no_irreducible_factors {x y : α} (nonzero : ¬(x = 0 ∧ y = 0)) - (H : ∀ z : α, Irreducible z → z ∣ x → ¬z ∣ y) : IsRelPrime x y := - isRelPrime_of_no_nonunits_factors nonzero fun _z znu znz zx zy ↦ - have ⟨i, h1, h2⟩ := exists_irreducible_factor znu znz - H i h1 (h2.trans zx) (h2.trans zy) - -end WfDvdMonoid - -theorem WfDvdMonoid.of_wellFoundedLT_associates [CancelCommMonoidWithZero α] - (h : WellFoundedLT (Associates α)) : WfDvdMonoid α := - WfDvdMonoid.of_wfDvdMonoid_associates - ⟨by - convert h.wf - ext - exact Associates.dvdNotUnit_iff_lt⟩ - -@[deprecated WfDvdMonoid.of_wellFoundedLT_associates (since := "2024-09-02")] -theorem WfDvdMonoid.of_wellFounded_associates [CancelCommMonoidWithZero α] - (h : WellFounded ((· < ·) : Associates α → Associates α → Prop)) : WfDvdMonoid α := - WfDvdMonoid.of_wfDvdMonoid_associates - ⟨by - convert h - ext - exact Associates.dvdNotUnit_iff_lt⟩ - -theorem WfDvdMonoid.iff_wellFounded_associates [CancelCommMonoidWithZero α] : - WfDvdMonoid α ↔ WellFoundedLT (Associates α) := - ⟨by apply WfDvdMonoid.wellFoundedLT_associates, WfDvdMonoid.of_wellFoundedLT_associates⟩ - -theorem WfDvdMonoid.max_power_factor' [CommMonoidWithZero α] [WfDvdMonoid α] {a₀ x : α} - (h : a₀ ≠ 0) (hx : ¬IsUnit x) : ∃ (n : ℕ) (a : α), ¬x ∣ a ∧ a₀ = x ^ n * a := by - obtain ⟨a, ⟨n, rfl⟩, hm⟩ := wellFounded_dvdNotUnit.has_min - {a | ∃ n, x ^ n * a = a₀} ⟨a₀, 0, by rw [pow_zero, one_mul]⟩ - refine ⟨n, a, ?_, rfl⟩; rintro ⟨d, rfl⟩ - exact hm d ⟨n + 1, by rw [pow_succ, mul_assoc]⟩ - ⟨(right_ne_zero_of_mul <| right_ne_zero_of_mul h), x, hx, mul_comm _ _⟩ - -theorem WfDvdMonoid.max_power_factor [CommMonoidWithZero α] [WfDvdMonoid α] {a₀ x : α} - (h : a₀ ≠ 0) (hx : Irreducible x) : ∃ (n : ℕ) (a : α), ¬x ∣ a ∧ a₀ = x ^ n * a := - max_power_factor' h hx.not_unit - -theorem multiplicity.finite_of_not_isUnit [CancelCommMonoidWithZero α] [WfDvdMonoid α] - {a b : α} (ha : ¬IsUnit a) (hb : b ≠ 0) : multiplicity.Finite a b := by - obtain ⟨n, c, ndvd, rfl⟩ := WfDvdMonoid.max_power_factor' hb ha - exact ⟨n, by rwa [pow_succ, mul_dvd_mul_iff_left (left_ne_zero_of_mul hb)]⟩ - -section Prio - --- set_option default_priority 100 - --- see Note [default priority] -/-- unique factorization monoids. - -These are defined as `CancelCommMonoidWithZero`s with well-founded strict divisibility -relations, but this is equivalent to more familiar definitions: - -Each element (except zero) is uniquely represented as a multiset of irreducible factors. -Uniqueness is only up to associated elements. - -Each element (except zero) is non-uniquely represented as a multiset -of prime factors. - -To define a UFD using the definition in terms of multisets -of irreducible factors, use the definition `of_exists_unique_irreducible_factors` - -To define a UFD using the definition in terms of multisets -of prime factors, use the definition `of_exists_prime_factors` - --/ -class UniqueFactorizationMonoid (α : Type*) [CancelCommMonoidWithZero α] extends - IsWellFounded α DvdNotUnit : Prop where - protected irreducible_iff_prime : ∀ {a : α}, Irreducible a ↔ Prime a - -instance (priority := 100) ufm_of_decomposition_of_wfDvdMonoid - [CancelCommMonoidWithZero α] [WfDvdMonoid α] [DecompositionMonoid α] : - UniqueFactorizationMonoid α := - { ‹WfDvdMonoid α› with irreducible_iff_prime := irreducible_iff_prime } - -@[deprecated ufm_of_decomposition_of_wfDvdMonoid (since := "2024-02-12")] -theorem ufm_of_gcd_of_wfDvdMonoid [CancelCommMonoidWithZero α] [WfDvdMonoid α] - [DecompositionMonoid α] : UniqueFactorizationMonoid α := - ufm_of_decomposition_of_wfDvdMonoid - -instance Associates.ufm [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] : - UniqueFactorizationMonoid (Associates α) := - { (WfDvdMonoid.wfDvdMonoid_associates : WfDvdMonoid (Associates α)) with - irreducible_iff_prime := by - rw [← Associates.irreducible_iff_prime_iff] - apply UniqueFactorizationMonoid.irreducible_iff_prime } - -end Prio - -namespace UniqueFactorizationMonoid - -variable [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] - -theorem exists_prime_factors (a : α) : - a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Prime b) ∧ f.prod ~ᵤ a := by - simp_rw [← UniqueFactorizationMonoid.irreducible_iff_prime] - apply WfDvdMonoid.exists_factors a - -instance : DecompositionMonoid α where - primal a := by - obtain rfl | ha := eq_or_ne a 0; · exact isPrimal_zero - obtain ⟨f, hf, u, rfl⟩ := exists_prime_factors a ha - exact ((Submonoid.isPrimal α).multiset_prod_mem f (hf · ·|>.isPrimal)).mul u.isUnit.isPrimal - -lemma exists_prime_iff : - (∃ (p : α), Prime p) ↔ ∃ (x : α), x ≠ 0 ∧ ¬ IsUnit x := by - refine ⟨fun ⟨p, hp⟩ ↦ ⟨p, hp.ne_zero, hp.not_unit⟩, fun ⟨x, hx₀, hxu⟩ ↦ ?_⟩ - obtain ⟨f, hf, -⟩ := WfDvdMonoid.exists_irreducible_factor hxu hx₀ - exact ⟨f, UniqueFactorizationMonoid.irreducible_iff_prime.mp hf⟩ - -@[elab_as_elim] -theorem induction_on_prime {P : α → Prop} (a : α) (h₁ : P 0) (h₂ : ∀ x : α, IsUnit x → P x) - (h₃ : ∀ a p : α, a ≠ 0 → Prime p → P a → P (p * a)) : P a := by - simp_rw [← UniqueFactorizationMonoid.irreducible_iff_prime] at h₃ - exact WfDvdMonoid.induction_on_irreducible a h₁ h₂ h₃ - -end UniqueFactorizationMonoid - -theorem prime_factors_unique [CancelCommMonoidWithZero α] : - ∀ {f g : Multiset α}, - (∀ x ∈ f, Prime x) → (∀ x ∈ g, Prime x) → f.prod ~ᵤ g.prod → Multiset.Rel Associated f g := by - classical - intro f - induction' f using Multiset.induction_on with p f ih - · intros g _ hg h - exact Multiset.rel_zero_left.2 <| - Multiset.eq_zero_of_forall_not_mem fun x hx => - have : IsUnit g.prod := by simpa [associated_one_iff_isUnit] using h.symm - (hg x hx).not_unit <| - isUnit_iff_dvd_one.2 <| (Multiset.dvd_prod hx).trans (isUnit_iff_dvd_one.1 this) - · intros g hf hg hfg - let ⟨b, hbg, hb⟩ := - (exists_associated_mem_of_dvd_prod (hf p (by simp)) fun q hq => hg _ hq) <| - hfg.dvd_iff_dvd_right.1 (show p ∣ (p ::ₘ f).prod by simp) - haveI := Classical.decEq α - rw [← Multiset.cons_erase hbg] - exact - Multiset.Rel.cons hb - (ih (fun q hq => hf _ (by simp [hq])) - (fun {q} (hq : q ∈ g.erase b) => hg q (Multiset.mem_of_mem_erase hq)) - (Associated.of_mul_left - (by rwa [← Multiset.prod_cons, ← Multiset.prod_cons, Multiset.cons_erase hbg]) hb - (hf p (by simp)).ne_zero)) - -namespace UniqueFactorizationMonoid - -variable [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] - -theorem factors_unique {f g : Multiset α} (hf : ∀ x ∈ f, Irreducible x) - (hg : ∀ x ∈ g, Irreducible x) (h : f.prod ~ᵤ g.prod) : Multiset.Rel Associated f g := - prime_factors_unique (fun x hx => UniqueFactorizationMonoid.irreducible_iff_prime.mp (hf x hx)) - (fun x hx => UniqueFactorizationMonoid.irreducible_iff_prime.mp (hg x hx)) h - -end UniqueFactorizationMonoid - -/-- If an irreducible has a prime factorization, - then it is an associate of one of its prime factors. -/ -theorem prime_factors_irreducible [CancelCommMonoidWithZero α] {a : α} {f : Multiset α} - (ha : Irreducible a) (pfa : (∀ b ∈ f, Prime b) ∧ f.prod ~ᵤ a) : ∃ p, a ~ᵤ p ∧ f = {p} := by - haveI := Classical.decEq α - refine @Multiset.induction_on _ - (fun g => (g.prod ~ᵤ a) → (∀ b ∈ g, Prime b) → ∃ p, a ~ᵤ p ∧ g = {p}) f ?_ ?_ pfa.2 pfa.1 - · intro h; exact (ha.not_unit (associated_one_iff_isUnit.1 (Associated.symm h))).elim - · rintro p s _ ⟨u, hu⟩ hs - use p - have hs0 : s = 0 := by - by_contra hs0 - obtain ⟨q, hq⟩ := Multiset.exists_mem_of_ne_zero hs0 - apply (hs q (by simp [hq])).2.1 - refine (ha.isUnit_or_isUnit (?_ : _ = p * ↑u * (s.erase q).prod * _)).resolve_left ?_ - · rw [mul_right_comm _ _ q, mul_assoc, ← Multiset.prod_cons, Multiset.cons_erase hq, ← hu, - mul_comm, mul_comm p _, mul_assoc] - simp - apply mt isUnit_of_mul_isUnit_left (mt isUnit_of_mul_isUnit_left _) - apply (hs p (Multiset.mem_cons_self _ _)).2.1 - simp only [mul_one, Multiset.prod_cons, Multiset.prod_zero, hs0] at * - exact ⟨Associated.symm ⟨u, hu⟩, rfl⟩ - -section ExistsPrimeFactors - -variable [CancelCommMonoidWithZero α] -variable (pf : ∀ a : α, a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Prime b) ∧ f.prod ~ᵤ a) -include pf - -theorem WfDvdMonoid.of_exists_prime_factors : WfDvdMonoid α := - ⟨by - classical - refine RelHomClass.wellFounded - (RelHom.mk ?_ ?_ : (DvdNotUnit : α → α → Prop) →r ((· < ·) : ℕ∞ → ℕ∞ → Prop)) wellFounded_lt - · intro a - by_cases h : a = 0 - · exact ⊤ - exact ↑(Multiset.card (Classical.choose (pf a h))) - rintro a b ⟨ane0, ⟨c, hc, b_eq⟩⟩ - rw [dif_neg ane0] - by_cases h : b = 0 - · simp [h, lt_top_iff_ne_top] - · rw [dif_neg h, Nat.cast_lt] - have cne0 : c ≠ 0 := by - refine mt (fun con => ?_) h - rw [b_eq, con, mul_zero] - calc - Multiset.card (Classical.choose (pf a ane0)) < - _ + Multiset.card (Classical.choose (pf c cne0)) := - lt_add_of_pos_right _ - (Multiset.card_pos.mpr fun con => hc (associated_one_iff_isUnit.mp ?_)) - _ = Multiset.card (Classical.choose (pf a ane0) + Classical.choose (pf c cne0)) := - (Multiset.card_add _ _).symm - _ = Multiset.card (Classical.choose (pf b h)) := - Multiset.card_eq_card_of_rel - (prime_factors_unique ?_ (Classical.choose_spec (pf _ h)).1 ?_) - - · convert (Classical.choose_spec (pf c cne0)).2.symm - rw [con, Multiset.prod_zero] - · intro x hadd - rw [Multiset.mem_add] at hadd - cases' hadd with h h <;> apply (Classical.choose_spec (pf _ _)).1 _ h <;> assumption - · rw [Multiset.prod_add] - trans a * c - · apply Associated.mul_mul <;> apply (Classical.choose_spec (pf _ _)).2 <;> assumption - · rw [← b_eq] - apply (Classical.choose_spec (pf _ _)).2.symm; assumption⟩ - -theorem irreducible_iff_prime_of_exists_prime_factors {p : α} : Irreducible p ↔ Prime p := by - by_cases hp0 : p = 0 - · simp [hp0] - refine ⟨fun h => ?_, Prime.irreducible⟩ - obtain ⟨f, hf⟩ := pf p hp0 - obtain ⟨q, hq, rfl⟩ := prime_factors_irreducible h hf - rw [hq.prime_iff] - exact hf.1 q (Multiset.mem_singleton_self _) - -theorem UniqueFactorizationMonoid.of_exists_prime_factors : UniqueFactorizationMonoid α := - { WfDvdMonoid.of_exists_prime_factors pf with - irreducible_iff_prime := irreducible_iff_prime_of_exists_prime_factors pf } - -end ExistsPrimeFactors - -theorem UniqueFactorizationMonoid.iff_exists_prime_factors [CancelCommMonoidWithZero α] : - UniqueFactorizationMonoid α ↔ - ∀ a : α, a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Prime b) ∧ f.prod ~ᵤ a := - ⟨fun h => @UniqueFactorizationMonoid.exists_prime_factors _ _ h, - UniqueFactorizationMonoid.of_exists_prime_factors⟩ - -section - -variable {β : Type*} [CancelCommMonoidWithZero α] [CancelCommMonoidWithZero β] - -theorem MulEquiv.uniqueFactorizationMonoid (e : α ≃* β) (hα : UniqueFactorizationMonoid α) : - UniqueFactorizationMonoid β := by - rw [UniqueFactorizationMonoid.iff_exists_prime_factors] at hα ⊢ - intro a ha - obtain ⟨w, hp, u, h⟩ := - hα (e.symm a) fun h => - ha <| by - convert← map_zero e - simp [← h] - exact - ⟨w.map e, fun b hb => - let ⟨c, hc, he⟩ := Multiset.mem_map.1 hb - he ▸ e.prime_iff.1 (hp c hc), - Units.map e.toMonoidHom u, - by - 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 β := - ⟨e.uniqueFactorizationMonoid, e.symm.uniqueFactorizationMonoid⟩ - -end - -theorem irreducible_iff_prime_of_exists_unique_irreducible_factors [CancelCommMonoidWithZero α] - (eif : ∀ a : α, a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Irreducible b) ∧ f.prod ~ᵤ a) - (uif : - ∀ f g : Multiset α, - (∀ x ∈ f, Irreducible x) → - (∀ x ∈ g, Irreducible x) → f.prod ~ᵤ g.prod → Multiset.Rel Associated f g) - (p : α) : Irreducible p ↔ Prime p := - letI := Classical.decEq α - ⟨ fun hpi => - ⟨hpi.ne_zero, hpi.1, fun a b ⟨x, hx⟩ => - if hab0 : a * b = 0 then - (eq_zero_or_eq_zero_of_mul_eq_zero hab0).elim (fun ha0 => by simp [ha0]) fun hb0 => by - simp [hb0] - else by - have hx0 : x ≠ 0 := fun hx0 => by simp_all - have ha0 : a ≠ 0 := left_ne_zero_of_mul hab0 - have hb0 : b ≠ 0 := right_ne_zero_of_mul hab0 - cases' eif x hx0 with fx hfx - cases' eif a ha0 with fa hfa - cases' eif b hb0 with fb hfb - have h : Multiset.Rel Associated (p ::ₘ fx) (fa + fb) := by - apply uif - · exact fun i hi => (Multiset.mem_cons.1 hi).elim (fun hip => hip.symm ▸ hpi) (hfx.1 _) - · exact fun i hi => (Multiset.mem_add.1 hi).elim (hfa.1 _) (hfb.1 _) - calc - Multiset.prod (p ::ₘ fx) ~ᵤ a * b := by - rw [hx, Multiset.prod_cons]; exact hfx.2.mul_left _ - _ ~ᵤ fa.prod * fb.prod := hfa.2.symm.mul_mul hfb.2.symm - _ = _ := by rw [Multiset.prod_add] - - exact - let ⟨q, hqf, hq⟩ := Multiset.exists_mem_of_rel_of_mem h (Multiset.mem_cons_self p _) - (Multiset.mem_add.1 hqf).elim - (fun hqa => - Or.inl <| hq.dvd_iff_dvd_left.2 <| hfa.2.dvd_iff_dvd_right.1 (Multiset.dvd_prod hqa)) - fun hqb => - Or.inr <| hq.dvd_iff_dvd_left.2 <| hfb.2.dvd_iff_dvd_right.1 (Multiset.dvd_prod hqb)⟩, - Prime.irreducible⟩ - -theorem UniqueFactorizationMonoid.of_exists_unique_irreducible_factors [CancelCommMonoidWithZero α] - (eif : ∀ a : α, a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Irreducible b) ∧ f.prod ~ᵤ a) - (uif : - ∀ f g : Multiset α, - (∀ x ∈ f, Irreducible x) → - (∀ x ∈ g, Irreducible x) → f.prod ~ᵤ g.prod → Multiset.Rel Associated f g) : - UniqueFactorizationMonoid α := - UniqueFactorizationMonoid.of_exists_prime_factors - (by - convert eif using 7 - simp_rw [irreducible_iff_prime_of_exists_unique_irreducible_factors eif uif]) - -namespace UniqueFactorizationMonoid - -variable [CancelCommMonoidWithZero α] -variable [UniqueFactorizationMonoid α] - -open Classical in -/-- Noncomputably determines the multiset of prime factors. -/ -noncomputable def factors (a : α) : Multiset α := - if h : a = 0 then 0 else Classical.choose (UniqueFactorizationMonoid.exists_prime_factors a h) - -theorem factors_prod {a : α} (ane0 : a ≠ 0) : Associated (factors a).prod a := by - rw [factors, dif_neg ane0] - exact (Classical.choose_spec (exists_prime_factors a ane0)).2 - -@[simp] -theorem factors_zero : factors (0 : α) = 0 := by simp [factors] - -theorem ne_zero_of_mem_factors {p a : α} (h : p ∈ factors a) : a ≠ 0 := by - rintro rfl - simp at h - -theorem dvd_of_mem_factors {p a : α} (h : p ∈ factors a) : p ∣ a := - dvd_trans (Multiset.dvd_prod h) (Associated.dvd (factors_prod (ne_zero_of_mem_factors h))) - -theorem prime_of_factor {a : α} (x : α) (hx : x ∈ factors a) : Prime x := by - have ane0 := ne_zero_of_mem_factors hx - rw [factors, dif_neg ane0] at hx - exact (Classical.choose_spec (UniqueFactorizationMonoid.exists_prime_factors a ane0)).1 x hx - -theorem irreducible_of_factor {a : α} : ∀ x : α, x ∈ factors a → Irreducible x := fun x h => - (prime_of_factor x h).irreducible - -@[simp] -theorem factors_one : factors (1 : α) = 0 := by - nontriviality α using factors - rw [← Multiset.rel_zero_right] - refine factors_unique irreducible_of_factor (fun x hx => (Multiset.not_mem_zero x hx).elim) ?_ - rw [Multiset.prod_zero] - exact factors_prod one_ne_zero - -theorem exists_mem_factors_of_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) : - p ∣ a → ∃ q ∈ factors a, p ~ᵤ q := fun ⟨b, hb⟩ => - have hb0 : b ≠ 0 := fun hb0 => by simp_all - have : Multiset.Rel Associated (p ::ₘ factors b) (factors a) := - factors_unique - (fun _ hx => (Multiset.mem_cons.1 hx).elim (fun h => h.symm ▸ hp) (irreducible_of_factor _)) - irreducible_of_factor - (Associated.symm <| - calc - Multiset.prod (factors a) ~ᵤ a := factors_prod ha0 - _ = p * b := hb - _ ~ᵤ Multiset.prod (p ::ₘ factors b) := by - rw [Multiset.prod_cons]; exact (factors_prod hb0).symm.mul_left _ - ) - Multiset.exists_mem_of_rel_of_mem this (by simp) - -theorem exists_mem_factors {x : α} (hx : x ≠ 0) (h : ¬IsUnit x) : ∃ p, p ∈ factors x := by - obtain ⟨p', hp', hp'x⟩ := WfDvdMonoid.exists_irreducible_factor h hx - obtain ⟨p, hp, _⟩ := exists_mem_factors_of_dvd hx hp' hp'x - exact ⟨p, hp⟩ - -open Classical in -theorem factors_mul {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : - Multiset.Rel Associated (factors (x * y)) (factors x + factors y) := by - refine - factors_unique irreducible_of_factor - (fun a ha => - (Multiset.mem_add.mp ha).by_cases (irreducible_of_factor _) (irreducible_of_factor _)) - ((factors_prod (mul_ne_zero hx hy)).trans ?_) - rw [Multiset.prod_add] - exact (Associated.mul_mul (factors_prod hx) (factors_prod hy)).symm - -theorem factors_pow {x : α} (n : ℕ) : - Multiset.Rel Associated (factors (x ^ n)) (n • factors x) := by - match n with - | 0 => rw [zero_smul, pow_zero, factors_one, Multiset.rel_zero_right] - | n+1 => - by_cases h0 : x = 0 - · simp [h0, zero_pow n.succ_ne_zero, smul_zero] - · rw [pow_succ', succ_nsmul'] - refine Multiset.Rel.trans _ (factors_mul h0 (pow_ne_zero n h0)) ?_ - refine Multiset.Rel.add ?_ <| factors_pow n - exact Multiset.rel_refl_of_refl_on fun y _ => Associated.refl _ - -@[simp] -theorem factors_pos (x : α) (hx : x ≠ 0) : 0 < factors x ↔ ¬IsUnit x := by - constructor - · intro h hx - obtain ⟨p, hp⟩ := Multiset.exists_mem_of_ne_zero h.ne' - exact (prime_of_factor _ hp).not_unit (isUnit_of_dvd_unit (dvd_of_mem_factors hp) hx) - · intro h - obtain ⟨p, hp⟩ := exists_mem_factors hx h - exact - bot_lt_iff_ne_bot.mpr - (mt Multiset.eq_zero_iff_forall_not_mem.mp (not_forall.mpr ⟨p, not_not.mpr hp⟩)) - -open Multiset in -theorem factors_pow_count_prod [DecidableEq α] {x : α} (hx : x ≠ 0) : - (∏ p ∈ (factors x).toFinset, p ^ (factors x).count p) ~ᵤ x := - calc - _ = prod (∑ a ∈ toFinset (factors x), count a (factors x) • {a}) := by - simp only [prod_sum, prod_nsmul, prod_singleton] - _ = prod (factors x) := by rw [toFinset_sum_count_nsmul_eq (factors x)] - _ ~ᵤ x := factors_prod hx - -theorem factors_rel_of_associated {a b : α} (h : Associated a b) : - Multiset.Rel Associated (factors a) (factors b) := by - rcases iff_iff_and_or_not_and_not.mp h.eq_zero_iff with (⟨rfl, rfl⟩ | ⟨ha, hb⟩) - · simp - · refine factors_unique irreducible_of_factor irreducible_of_factor ?_ - exact ((factors_prod ha).trans h).trans (factors_prod hb).symm - -end UniqueFactorizationMonoid - -namespace UniqueFactorizationMonoid - -variable [CancelCommMonoidWithZero α] [NormalizationMonoid α] -variable [UniqueFactorizationMonoid α] - -/-- Noncomputably determines the multiset of prime factors. -/ -noncomputable def normalizedFactors (a : α) : Multiset α := - Multiset.map normalize <| factors a - -/-- An arbitrary choice of factors of `x : M` is exactly the (unique) normalized set of factors, -if `M` has a trivial group of units. -/ -@[simp] -theorem factors_eq_normalizedFactors {M : Type*} [CancelCommMonoidWithZero M] - [UniqueFactorizationMonoid M] [Subsingleton Mˣ] (x : M) : factors x = normalizedFactors x := by - unfold normalizedFactors - convert (Multiset.map_id (factors x)).symm - ext p - exact normalize_eq p - -theorem normalizedFactors_prod {a : α} (ane0 : a ≠ 0) : - Associated (normalizedFactors a).prod a := by - rw [normalizedFactors, factors, dif_neg ane0] - refine Associated.trans ?_ (Classical.choose_spec (exists_prime_factors a ane0)).2 - rw [← Associates.mk_eq_mk_iff_associated, ← Associates.prod_mk, ← Associates.prod_mk, - Multiset.map_map] - congr 2 - ext - rw [Function.comp_apply, Associates.mk_normalize] - -theorem prime_of_normalized_factor {a : α} : ∀ x : α, x ∈ normalizedFactors a → Prime x := by - rw [normalizedFactors, factors] - split_ifs with ane0; · simp - intro x hx; rcases Multiset.mem_map.1 hx with ⟨y, ⟨hy, rfl⟩⟩ - rw [(normalize_associated _).prime_iff] - exact (Classical.choose_spec (UniqueFactorizationMonoid.exists_prime_factors a ane0)).1 y hy - -theorem irreducible_of_normalized_factor {a : α} : - ∀ x : α, x ∈ normalizedFactors a → Irreducible x := fun x h => - (prime_of_normalized_factor x h).irreducible - -theorem normalize_normalized_factor {a : α} : - ∀ x : α, x ∈ normalizedFactors a → normalize x = x := by - rw [normalizedFactors, factors] - split_ifs with h; · simp - intro x hx - obtain ⟨y, _, rfl⟩ := Multiset.mem_map.1 hx - apply normalize_idem - -theorem normalizedFactors_irreducible {a : α} (ha : Irreducible a) : - normalizedFactors a = {normalize a} := by - obtain ⟨p, a_assoc, hp⟩ := - prime_factors_irreducible ha ⟨prime_of_normalized_factor, normalizedFactors_prod ha.ne_zero⟩ - have p_mem : p ∈ normalizedFactors a := by - rw [hp] - exact Multiset.mem_singleton_self _ - convert hp - rwa [← normalize_normalized_factor p p_mem, normalize_eq_normalize_iff, dvd_dvd_iff_associated] - -theorem normalizedFactors_eq_of_dvd (a : α) : - ∀ᵉ (p ∈ normalizedFactors a) (q ∈ normalizedFactors a), p ∣ q → p = q := by - intro p hp q hq hdvd - convert normalize_eq_normalize hdvd - ((prime_of_normalized_factor _ hp).irreducible.dvd_symm - (prime_of_normalized_factor _ hq).irreducible hdvd) <;> - apply (normalize_normalized_factor _ ‹_›).symm - -theorem exists_mem_normalizedFactors_of_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) : - p ∣ a → ∃ q ∈ normalizedFactors a, p ~ᵤ q := fun ⟨b, hb⟩ => - have hb0 : b ≠ 0 := fun hb0 => by simp_all - have : Multiset.Rel Associated (p ::ₘ normalizedFactors b) (normalizedFactors a) := - factors_unique - (fun _ hx => - (Multiset.mem_cons.1 hx).elim (fun h => h.symm ▸ hp) (irreducible_of_normalized_factor _)) - irreducible_of_normalized_factor - (Associated.symm <| - calc - Multiset.prod (normalizedFactors a) ~ᵤ a := normalizedFactors_prod ha0 - _ = p * b := hb - _ ~ᵤ Multiset.prod (p ::ₘ normalizedFactors b) := by - rw [Multiset.prod_cons] - exact (normalizedFactors_prod hb0).symm.mul_left _ - ) - Multiset.exists_mem_of_rel_of_mem this (by simp) - -theorem exists_mem_normalizedFactors {x : α} (hx : x ≠ 0) (h : ¬IsUnit x) : - ∃ p, p ∈ normalizedFactors x := by - obtain ⟨p', hp', hp'x⟩ := WfDvdMonoid.exists_irreducible_factor h hx - obtain ⟨p, hp, _⟩ := exists_mem_normalizedFactors_of_dvd hx hp' hp'x - exact ⟨p, hp⟩ - -@[simp] -theorem normalizedFactors_zero : normalizedFactors (0 : α) = 0 := by - simp [normalizedFactors, factors] - -@[simp] -theorem normalizedFactors_one : normalizedFactors (1 : α) = 0 := by - cases' subsingleton_or_nontrivial α with h h - · dsimp [normalizedFactors, factors] - simp [Subsingleton.elim (1 : α) 0] - · rw [← Multiset.rel_zero_right] - apply factors_unique irreducible_of_normalized_factor - · intro x hx - exfalso - apply Multiset.not_mem_zero x hx - · apply normalizedFactors_prod one_ne_zero - -@[simp] -theorem normalizedFactors_mul {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : - normalizedFactors (x * y) = normalizedFactors x + normalizedFactors y := by - have h : (normalize : α → α) = Associates.out ∘ Associates.mk := by - ext - rw [Function.comp_apply, Associates.out_mk] - rw [← Multiset.map_id' (normalizedFactors (x * y)), ← Multiset.map_id' (normalizedFactors x), ← - Multiset.map_id' (normalizedFactors y), ← Multiset.map_congr rfl normalize_normalized_factor, ← - Multiset.map_congr rfl normalize_normalized_factor, ← - Multiset.map_congr rfl normalize_normalized_factor, ← Multiset.map_add, h, ← - Multiset.map_map Associates.out, eq_comm, ← Multiset.map_map Associates.out] - refine congr rfl ?_ - apply Multiset.map_mk_eq_map_mk_of_rel - apply factors_unique - · intro x hx - rcases Multiset.mem_add.1 hx with (hx | hx) <;> exact irreducible_of_normalized_factor x hx - · exact irreducible_of_normalized_factor - · rw [Multiset.prod_add] - exact - ((normalizedFactors_prod hx).mul_mul (normalizedFactors_prod hy)).trans - (normalizedFactors_prod (mul_ne_zero hx hy)).symm - -@[simp] -theorem normalizedFactors_pow {x : α} (n : ℕ) : - normalizedFactors (x ^ n) = n • normalizedFactors x := by - induction' n with n ih - · simp - by_cases h0 : x = 0 - · simp [h0, zero_pow n.succ_ne_zero, smul_zero] - rw [pow_succ', succ_nsmul', normalizedFactors_mul h0 (pow_ne_zero _ h0), ih] - -theorem _root_.Irreducible.normalizedFactors_pow {p : α} (hp : Irreducible p) (k : ℕ) : - normalizedFactors (p ^ k) = Multiset.replicate k (normalize p) := by - rw [UniqueFactorizationMonoid.normalizedFactors_pow, normalizedFactors_irreducible hp, - Multiset.nsmul_singleton] - -theorem normalizedFactors_prod_eq (s : Multiset α) (hs : ∀ a ∈ s, Irreducible a) : - normalizedFactors s.prod = s.map normalize := by - induction' s using Multiset.induction with a s ih - · rw [Multiset.prod_zero, normalizedFactors_one, Multiset.map_zero] - · have ia := hs a (Multiset.mem_cons_self a _) - have ib := fun b h => hs b (Multiset.mem_cons_of_mem h) - obtain rfl | ⟨b, hb⟩ := s.empty_or_exists_mem - · rw [Multiset.cons_zero, Multiset.prod_singleton, Multiset.map_singleton, - normalizedFactors_irreducible ia] - haveI := nontrivial_of_ne b 0 (ib b hb).ne_zero - rw [Multiset.prod_cons, Multiset.map_cons, - normalizedFactors_mul ia.ne_zero (Multiset.prod_ne_zero fun h => (ib 0 h).ne_zero rfl), - normalizedFactors_irreducible ia, ih ib, Multiset.singleton_add] - -theorem dvd_iff_normalizedFactors_le_normalizedFactors {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : - x ∣ y ↔ normalizedFactors x ≤ normalizedFactors y := by - constructor - · rintro ⟨c, rfl⟩ - simp [hx, right_ne_zero_of_mul hy] - · rw [← (normalizedFactors_prod hx).dvd_iff_dvd_left, ← - (normalizedFactors_prod hy).dvd_iff_dvd_right] - apply Multiset.prod_dvd_prod_of_le - -theorem _root_.Associated.normalizedFactors_eq {a b : α} (h : Associated a b) : - normalizedFactors a = normalizedFactors b := by - unfold normalizedFactors - have h' : ⇑(normalize (α := α)) = Associates.out ∘ Associates.mk := funext Associates.out_mk - rw [h', ← Multiset.map_map, ← Multiset.map_map, - Associates.rel_associated_iff_map_eq_map.mp (factors_rel_of_associated h)] - -theorem associated_iff_normalizedFactors_eq_normalizedFactors {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : - x ~ᵤ y ↔ normalizedFactors x = normalizedFactors y := - ⟨Associated.normalizedFactors_eq, fun h => - (normalizedFactors_prod hx).symm.trans (_root_.trans (by rw [h]) (normalizedFactors_prod hy))⟩ - -theorem normalizedFactors_of_irreducible_pow {p : α} (hp : Irreducible p) (k : ℕ) : - normalizedFactors (p ^ k) = Multiset.replicate k (normalize p) := by - rw [normalizedFactors_pow, normalizedFactors_irreducible hp, Multiset.nsmul_singleton] - -theorem zero_not_mem_normalizedFactors (x : α) : (0 : α) ∉ normalizedFactors x := fun h => - Prime.ne_zero (prime_of_normalized_factor _ h) rfl - -theorem dvd_of_mem_normalizedFactors {a p : α} (H : p ∈ normalizedFactors a) : p ∣ a := by - by_cases hcases : a = 0 - · rw [hcases] - exact dvd_zero p - · exact dvd_trans (Multiset.dvd_prod H) (Associated.dvd (normalizedFactors_prod hcases)) - -theorem mem_normalizedFactors_iff [Subsingleton αˣ] {p x : α} (hx : x ≠ 0) : - p ∈ normalizedFactors x ↔ Prime p ∧ p ∣ x := by - constructor - · intro h - exact ⟨prime_of_normalized_factor p h, dvd_of_mem_normalizedFactors h⟩ - · rintro ⟨hprime, hdvd⟩ - obtain ⟨q, hqmem, hqeq⟩ := exists_mem_normalizedFactors_of_dvd hx hprime.irreducible hdvd - rw [associated_iff_eq] at hqeq - exact hqeq ▸ hqmem - -theorem exists_associated_prime_pow_of_unique_normalized_factor {p r : α} - (h : ∀ {m}, m ∈ normalizedFactors r → m = p) (hr : r ≠ 0) : ∃ i : ℕ, Associated (p ^ i) r := by - use Multiset.card.toFun (normalizedFactors r) - have := UniqueFactorizationMonoid.normalizedFactors_prod hr - rwa [Multiset.eq_replicate_of_mem fun b => h, Multiset.prod_replicate] at this - -theorem normalizedFactors_prod_of_prime [Subsingleton αˣ] {m : Multiset α} - (h : ∀ p ∈ m, Prime p) : normalizedFactors m.prod = m := by - cases subsingleton_or_nontrivial α - · obtain rfl : m = 0 := by - refine Multiset.eq_zero_of_forall_not_mem fun x hx ↦ ?_ - simpa [Subsingleton.elim x 0] using h x hx - simp - · simpa only [← Multiset.rel_eq, ← associated_eq_eq] using - prime_factors_unique prime_of_normalized_factor h - (normalizedFactors_prod (m.prod_ne_zero_of_prime h)) - -theorem mem_normalizedFactors_eq_of_associated {a b c : α} (ha : a ∈ normalizedFactors c) - (hb : b ∈ normalizedFactors c) (h : Associated a b) : a = b := by - rw [← normalize_normalized_factor a ha, ← normalize_normalized_factor b hb, - normalize_eq_normalize_iff] - exact Associated.dvd_dvd h - -@[simp] -theorem normalizedFactors_pos (x : α) (hx : x ≠ 0) : 0 < normalizedFactors x ↔ ¬IsUnit x := by - constructor - · intro h hx - obtain ⟨p, hp⟩ := Multiset.exists_mem_of_ne_zero h.ne' - exact - (prime_of_normalized_factor _ hp).not_unit - (isUnit_of_dvd_unit (dvd_of_mem_normalizedFactors hp) hx) - · intro h - obtain ⟨p, hp⟩ := exists_mem_normalizedFactors hx h - exact - bot_lt_iff_ne_bot.mpr - (mt Multiset.eq_zero_iff_forall_not_mem.mp (not_forall.mpr ⟨p, not_not.mpr hp⟩)) - -theorem dvdNotUnit_iff_normalizedFactors_lt_normalizedFactors {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : - DvdNotUnit x y ↔ normalizedFactors x < normalizedFactors y := by - constructor - · rintro ⟨_, c, hc, rfl⟩ - simp only [hx, right_ne_zero_of_mul hy, normalizedFactors_mul, Ne, not_false_iff, - lt_add_iff_pos_right, normalizedFactors_pos, hc] - · intro h - exact - dvdNotUnit_of_dvd_of_not_dvd - ((dvd_iff_normalizedFactors_le_normalizedFactors hx hy).mpr h.le) - (mt (dvd_iff_normalizedFactors_le_normalizedFactors hy hx).mp h.not_le) - -theorem normalizedFactors_multiset_prod (s : Multiset α) (hs : 0 ∉ s) : - normalizedFactors (s.prod) = (s.map normalizedFactors).sum := by - cases subsingleton_or_nontrivial α - · obtain rfl : s = 0 := by - apply Multiset.eq_zero_of_forall_not_mem - intro _ - convert hs - simp - induction s using Multiset.induction with - | empty => simp - | cons _ _ IH => - rw [Multiset.prod_cons, Multiset.map_cons, Multiset.sum_cons, normalizedFactors_mul, IH] - · exact fun h ↦ hs (Multiset.mem_cons_of_mem h) - · exact fun h ↦ hs (h ▸ Multiset.mem_cons_self _ _) - · apply Multiset.prod_ne_zero - exact fun h ↦ hs (Multiset.mem_cons_of_mem h) - -end UniqueFactorizationMonoid - -namespace UniqueFactorizationMonoid - -open scoped Classical - -open Multiset Associates - -variable [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] - -/-- Noncomputably defines a `normalizationMonoid` structure on a `UniqueFactorizationMonoid`. -/ -protected noncomputable def normalizationMonoid : NormalizationMonoid α := - normalizationMonoidOfMonoidHomRightInverse - { toFun := fun a : Associates α => - if a = 0 then 0 - else - ((normalizedFactors a).map - (Classical.choose mk_surjective.hasRightInverse : Associates α → α)).prod - map_one' := by nontriviality α; simp - map_mul' := fun x y => by - by_cases hx : x = 0 - · simp [hx] - by_cases hy : y = 0 - · simp [hy] - simp [hx, hy] } - (by - intro x - dsimp - by_cases hx : x = 0 - · simp [hx] - have h : Associates.mkMonoidHom ∘ Classical.choose mk_surjective.hasRightInverse = - (id : Associates α → Associates α) := by - ext x - rw [Function.comp_apply, mkMonoidHom_apply, - Classical.choose_spec mk_surjective.hasRightInverse x] - rfl - rw [if_neg hx, ← mkMonoidHom_apply, MonoidHom.map_multiset_prod, map_map, h, map_id, ← - associated_iff_eq] - apply normalizedFactors_prod hx) - -end UniqueFactorizationMonoid - -namespace UniqueFactorizationMonoid - -variable {R : Type*} [CancelCommMonoidWithZero R] [UniqueFactorizationMonoid R] - -theorem isRelPrime_iff_no_prime_factors {a b : R} (ha : a ≠ 0) : - IsRelPrime a b ↔ ∀ ⦃d⦄, d ∣ a → d ∣ b → ¬Prime d := - ⟨fun h _ ha hb ↦ (·.not_unit <| h ha hb), fun h ↦ WfDvdMonoid.isRelPrime_of_no_irreducible_factors - (ha ·.1) fun _ irr ha hb ↦ h ha hb (UniqueFactorizationMonoid.irreducible_iff_prime.mp irr)⟩ - -/-- Euclid's lemma: if `a ∣ b * c` and `a` and `c` have no common prime factors, `a ∣ b`. -Compare `IsCoprime.dvd_of_dvd_mul_left`. -/ -theorem dvd_of_dvd_mul_left_of_no_prime_factors {a b c : R} (ha : a ≠ 0) - (h : ∀ ⦃d⦄, d ∣ a → d ∣ c → ¬Prime d) : a ∣ b * c → a ∣ b := - ((isRelPrime_iff_no_prime_factors ha).mpr h).dvd_of_dvd_mul_right - -/-- Euclid's lemma: if `a ∣ b * c` and `a` and `b` have no common prime factors, `a ∣ c`. -Compare `IsCoprime.dvd_of_dvd_mul_right`. -/ -theorem dvd_of_dvd_mul_right_of_no_prime_factors {a b c : R} (ha : a ≠ 0) - (no_factors : ∀ {d}, d ∣ a → d ∣ b → ¬Prime d) : a ∣ b * c → a ∣ c := by - simpa [mul_comm b c] using dvd_of_dvd_mul_left_of_no_prime_factors ha @no_factors - -/-- If `a ≠ 0, b` are elements of a unique factorization domain, then dividing -out their common factor `c'` gives `a'` and `b'` with no factors in common. -/ -theorem exists_reduced_factors : - ∀ a ≠ (0 : R), ∀ b, - ∃ a' b' c', IsRelPrime a' b' ∧ c' * a' = a ∧ c' * b' = b := by - intro a - refine induction_on_prime a ?_ ?_ ?_ - · intros - contradiction - · intro a a_unit _ b - use a, b, 1 - constructor - · intro p p_dvd_a _ - exact isUnit_of_dvd_unit p_dvd_a a_unit - · simp - · intro a p a_ne_zero p_prime ih_a pa_ne_zero b - by_cases h : p ∣ b - · rcases h with ⟨b, rfl⟩ - obtain ⟨a', b', c', no_factor, ha', hb'⟩ := ih_a a_ne_zero b - refine ⟨a', b', p * c', @no_factor, ?_, ?_⟩ - · rw [mul_assoc, ha'] - · rw [mul_assoc, hb'] - · obtain ⟨a', b', c', coprime, rfl, rfl⟩ := ih_a a_ne_zero b - refine ⟨p * a', b', c', ?_, mul_left_comm _ _ _, rfl⟩ - intro q q_dvd_pa' q_dvd_b' - cases' p_prime.left_dvd_or_dvd_right_of_dvd_mul q_dvd_pa' with p_dvd_q q_dvd_a' - · have : p ∣ c' * b' := dvd_mul_of_dvd_right (p_dvd_q.trans q_dvd_b') _ - contradiction - exact coprime q_dvd_a' q_dvd_b' - -theorem exists_reduced_factors' (a b : R) (hb : b ≠ 0) : - ∃ a' b' c', IsRelPrime a' b' ∧ c' * a' = a ∧ c' * b' = b := - let ⟨b', a', c', no_factor, hb, ha⟩ := exists_reduced_factors b hb a - ⟨a', b', c', fun _ hpb hpa => no_factor hpa hpb, ha, hb⟩ - -@[deprecated (since := "2024-09-21")] alias pow_right_injective := pow_injective_of_not_isUnit -@[deprecated (since := "2024-09-21")] alias pow_eq_pow_iff := pow_inj_of_not_isUnit - -section multiplicity - -variable [NormalizationMonoid R] - -open Multiset - -section - -theorem le_emultiplicity_iff_replicate_le_normalizedFactors {a b : R} {n : ℕ} (ha : Irreducible a) - (hb : b ≠ 0) : - ↑n ≤ emultiplicity a b ↔ replicate n (normalize a) ≤ normalizedFactors b := by - rw [← pow_dvd_iff_le_emultiplicity] - revert b - induction' n with n ih; · simp - intro b hb - constructor - · rintro ⟨c, rfl⟩ - rw [Ne, pow_succ', mul_assoc, mul_eq_zero, not_or] at hb - rw [pow_succ', mul_assoc, normalizedFactors_mul hb.1 hb.2, replicate_succ, - normalizedFactors_irreducible ha, singleton_add, cons_le_cons_iff, ← ih hb.2] - apply Dvd.intro _ rfl - · rw [Multiset.le_iff_exists_add] - rintro ⟨u, hu⟩ - rw [← (normalizedFactors_prod hb).dvd_iff_dvd_right, hu, prod_add, prod_replicate] - exact (Associated.pow_pow <| associated_normalize a).dvd.trans (Dvd.intro u.prod rfl) - -/-- The multiplicity of an irreducible factor of a nonzero element is exactly the number of times -the normalized factor occurs in the `normalizedFactors`. - -See also `count_normalizedFactors_eq` which expands the definition of `multiplicity` -to produce a specification for `count (normalizedFactors _) _`.. --/ -theorem emultiplicity_eq_count_normalizedFactors [DecidableEq R] {a b : R} (ha : Irreducible a) - (hb : b ≠ 0) : emultiplicity a b = (normalizedFactors b).count (normalize a) := by - apply le_antisymm - · apply Order.le_of_lt_add_one - rw [← Nat.cast_one, ← Nat.cast_add, lt_iff_not_ge, ge_iff_le, - le_emultiplicity_iff_replicate_le_normalizedFactors ha hb, ← le_count_iff_replicate_le] - simp - rw [le_emultiplicity_iff_replicate_le_normalizedFactors ha hb, ← le_count_iff_replicate_le] - -end - -/-- The number of times an irreducible factor `p` appears in `normalizedFactors x` is defined by -the number of times it divides `x`. - -See also `multiplicity_eq_count_normalizedFactors` if `n` is given by `multiplicity p x`. --/ -theorem count_normalizedFactors_eq [DecidableEq R] {p x : R} (hp : Irreducible p) - (hnorm : normalize p = p) {n : ℕ} (hle : p ^ n ∣ x) (hlt : ¬p ^ (n + 1) ∣ x) : - (normalizedFactors x).count p = n := by classical - by_cases hx0 : x = 0 - · simp [hx0] at hlt - apply Nat.cast_injective (R := ℕ∞) - convert (emultiplicity_eq_count_normalizedFactors hp hx0).symm - · exact hnorm.symm - exact (emultiplicity_eq_coe.mpr ⟨hle, hlt⟩).symm - -/-- The number of times an irreducible factor `p` appears in `normalizedFactors x` is defined by -the number of times it divides `x`. This is a slightly more general version of -`UniqueFactorizationMonoid.count_normalizedFactors_eq` that allows `p = 0`. - -See also `multiplicity_eq_count_normalizedFactors` if `n` is given by `multiplicity p x`. --/ -theorem count_normalizedFactors_eq' [DecidableEq R] {p x : R} (hp : p = 0 ∨ Irreducible p) - (hnorm : normalize p = p) {n : ℕ} (hle : p ^ n ∣ x) (hlt : ¬p ^ (n + 1) ∣ x) : - (normalizedFactors x).count p = n := by - rcases hp with (rfl | hp) - · cases n - · exact count_eq_zero.2 (zero_not_mem_normalizedFactors _) - · rw [zero_pow (Nat.succ_ne_zero _)] at hle hlt - exact absurd hle hlt - · exact count_normalizedFactors_eq hp hnorm hle hlt - -end multiplicity - -/-- Deprecated. Use `WfDvdMonoid.max_power_factor` instead. -/ -@[deprecated WfDvdMonoid.max_power_factor (since := "2024-03-01")] -theorem max_power_factor {a₀ x : R} (h : a₀ ≠ 0) (hx : Irreducible x) : - ∃ n : ℕ, ∃ a : R, ¬x ∣ a ∧ a₀ = x ^ n * a := WfDvdMonoid.max_power_factor h hx - -section Multiplicative - -variable [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] -variable {β : Type*} [CancelCommMonoidWithZero β] - -theorem prime_pow_coprime_prod_of_coprime_insert [DecidableEq α] {s : Finset α} (i : α → ℕ) (p : α) - (hps : p ∉ s) (is_prime : ∀ q ∈ insert p s, Prime q) - (is_coprime : ∀ᵉ (q ∈ insert p s) (q' ∈ insert p s), q ∣ q' → q = q') : - IsRelPrime (p ^ i p) (∏ p' ∈ s, p' ^ i p') := by - have hp := is_prime _ (Finset.mem_insert_self _ _) - refine (isRelPrime_iff_no_prime_factors <| pow_ne_zero _ hp.ne_zero).mpr ?_ - intro d hdp hdprod hd - apply hps - replace hdp := hd.dvd_of_dvd_pow hdp - obtain ⟨q, q_mem', hdq⟩ := hd.exists_mem_multiset_dvd hdprod - obtain ⟨q, q_mem, rfl⟩ := Multiset.mem_map.mp q_mem' - replace hdq := hd.dvd_of_dvd_pow hdq - have : p ∣ q := dvd_trans (hd.irreducible.dvd_symm hp.irreducible hdp) hdq - convert q_mem using 0 - rw [Finset.mem_val, - is_coprime _ (Finset.mem_insert_self p s) _ (Finset.mem_insert_of_mem q_mem) this] - -/-- If `P` holds for units and powers of primes, -and `P x ∧ P y` for coprime `x, y` implies `P (x * y)`, -then `P` holds on a product of powers of distinct primes. -/ -@[elab_as_elim] -theorem induction_on_prime_power {P : α → Prop} (s : Finset α) (i : α → ℕ) - (is_prime : ∀ p ∈ s, Prime p) (is_coprime : ∀ᵉ (p ∈ s) (q ∈ s), p ∣ q → p = q) - (h1 : ∀ {x}, IsUnit x → P x) (hpr : ∀ {p} (i : ℕ), Prime p → P (p ^ i)) - (hcp : ∀ {x y}, IsRelPrime x y → P x → P y → P (x * y)) : - P (∏ p ∈ s, p ^ i p) := by - letI := Classical.decEq α - induction' s using Finset.induction_on with p f' hpf' ih - · simpa using h1 isUnit_one - rw [Finset.prod_insert hpf'] - exact - hcp (prime_pow_coprime_prod_of_coprime_insert i p hpf' is_prime is_coprime) - (hpr (i p) (is_prime _ (Finset.mem_insert_self _ _))) - (ih (fun q hq => is_prime _ (Finset.mem_insert_of_mem hq)) fun q hq q' hq' => - is_coprime _ (Finset.mem_insert_of_mem hq) _ (Finset.mem_insert_of_mem hq')) - -/-- If `P` holds for `0`, units and powers of primes, -and `P x ∧ P y` for coprime `x, y` implies `P (x * y)`, -then `P` holds on all `a : α`. -/ -@[elab_as_elim] -theorem induction_on_coprime {P : α → Prop} (a : α) (h0 : P 0) (h1 : ∀ {x}, IsUnit x → P x) - (hpr : ∀ {p} (i : ℕ), Prime p → P (p ^ i)) - (hcp : ∀ {x y}, IsRelPrime x y → P x → P y → P (x * y)) : P a := by - letI := Classical.decEq α - have P_of_associated : ∀ {x y}, Associated x y → P x → P y := by - rintro x y ⟨u, rfl⟩ hx - exact hcp (fun p _ hpx => isUnit_of_dvd_unit hpx u.isUnit) hx (h1 u.isUnit) - by_cases ha0 : a = 0 - · rwa [ha0] - haveI : Nontrivial α := ⟨⟨_, _, ha0⟩⟩ - letI : NormalizationMonoid α := UniqueFactorizationMonoid.normalizationMonoid - refine P_of_associated (normalizedFactors_prod ha0) ?_ - rw [← (normalizedFactors a).map_id, Finset.prod_multiset_map_count] - refine induction_on_prime_power _ _ ?_ ?_ @h1 @hpr @hcp <;> simp only [Multiset.mem_toFinset] - · apply prime_of_normalized_factor - · apply normalizedFactors_eq_of_dvd - -/-- If `f` maps `p ^ i` to `(f p) ^ i` for primes `p`, and `f` -is multiplicative on coprime elements, then `f` is multiplicative on all products of primes. -/ -theorem multiplicative_prime_power {f : α → β} (s : Finset α) (i j : α → ℕ) - (is_prime : ∀ p ∈ s, Prime p) (is_coprime : ∀ᵉ (p ∈ s) (q ∈ s), p ∣ q → p = q) - (h1 : ∀ {x y}, IsUnit y → f (x * y) = f x * f y) - (hpr : ∀ {p} (i : ℕ), Prime p → f (p ^ i) = f p ^ i) - (hcp : ∀ {x y}, IsRelPrime x y → f (x * y) = f x * f y) : - f (∏ p ∈ s, p ^ (i p + j p)) = f (∏ p ∈ s, p ^ i p) * f (∏ p ∈ s, p ^ j p) := by - letI := Classical.decEq α - induction' s using Finset.induction_on with p s hps ih - · simpa using h1 isUnit_one - have hpr_p := is_prime _ (Finset.mem_insert_self _ _) - have hpr_s : ∀ p ∈ s, Prime p := fun p hp => is_prime _ (Finset.mem_insert_of_mem hp) - have hcp_p := fun i => prime_pow_coprime_prod_of_coprime_insert i p hps is_prime is_coprime - have hcp_s : ∀ᵉ (p ∈ s) (q ∈ s), p ∣ q → p = q := fun p hp q hq => - is_coprime p (Finset.mem_insert_of_mem hp) q (Finset.mem_insert_of_mem hq) - rw [Finset.prod_insert hps, Finset.prod_insert hps, Finset.prod_insert hps, hcp (hcp_p _), - hpr _ hpr_p, hcp (hcp_p _), hpr _ hpr_p, hcp (hcp_p (fun p => i p + j p)), hpr _ hpr_p, - ih hpr_s hcp_s, pow_add, mul_assoc, mul_left_comm (f p ^ j p), mul_assoc] - -/-- If `f` maps `p ^ i` to `(f p) ^ i` for primes `p`, and `f` -is multiplicative on coprime elements, then `f` is multiplicative everywhere. -/ -theorem multiplicative_of_coprime (f : α → β) (a b : α) (h0 : f 0 = 0) - (h1 : ∀ {x y}, IsUnit y → f (x * y) = f x * f y) - (hpr : ∀ {p} (i : ℕ), Prime p → f (p ^ i) = f p ^ i) - (hcp : ∀ {x y}, IsRelPrime x y → f (x * y) = f x * f y) : - f (a * b) = f a * f b := by - letI := Classical.decEq α - by_cases ha0 : a = 0 - · rw [ha0, zero_mul, h0, zero_mul] - by_cases hb0 : b = 0 - · rw [hb0, mul_zero, h0, mul_zero] - by_cases hf1 : f 1 = 0 - · calc - f (a * b) = f (a * b * 1) := by rw [mul_one] - _ = 0 := by simp only [h1 isUnit_one, hf1, mul_zero] - _ = f a * f (b * 1) := by simp only [h1 isUnit_one, hf1, mul_zero] - _ = f a * f b := by rw [mul_one] - haveI : Nontrivial α := ⟨⟨_, _, ha0⟩⟩ - letI : NormalizationMonoid α := UniqueFactorizationMonoid.normalizationMonoid - suffices - f (∏ p ∈ (normalizedFactors a).toFinset ∪ (normalizedFactors b).toFinset, - p ^ ((normalizedFactors a).count p + (normalizedFactors b).count p)) = - f (∏ p ∈ (normalizedFactors a).toFinset ∪ (normalizedFactors b).toFinset, - p ^ (normalizedFactors a).count p) * - f (∏ p ∈ (normalizedFactors a).toFinset ∪ (normalizedFactors b).toFinset, - p ^ (normalizedFactors b).count p) by - obtain ⟨ua, a_eq⟩ := normalizedFactors_prod ha0 - obtain ⟨ub, b_eq⟩ := normalizedFactors_prod hb0 - rw [← a_eq, ← b_eq, mul_right_comm (Multiset.prod (normalizedFactors a)) ua - (Multiset.prod (normalizedFactors b) * ub), h1 ua.isUnit, h1 ub.isUnit, h1 ua.isUnit, ← - mul_assoc, h1 ub.isUnit, mul_right_comm _ (f ua), ← mul_assoc] - congr - rw [← (normalizedFactors a).map_id, ← (normalizedFactors b).map_id, - Finset.prod_multiset_map_count, Finset.prod_multiset_map_count, - Finset.prod_subset (Finset.subset_union_left (s₂ := (normalizedFactors b).toFinset)), - Finset.prod_subset (Finset.subset_union_right (s₂ := (normalizedFactors b).toFinset)), ← - Finset.prod_mul_distrib] - · simp_rw [id, ← pow_add, this] - all_goals simp only [Multiset.mem_toFinset] - · intro p _ hpb - simp [hpb] - · intro p _ hpa - simp [hpa] - refine multiplicative_prime_power _ _ _ ?_ ?_ @h1 @hpr @hcp - all_goals simp only [Multiset.mem_toFinset, Finset.mem_union] - · rintro p (hpa | hpb) <;> apply prime_of_normalized_factor <;> assumption - · rintro p (hp | hp) q (hq | hq) hdvd <;> - rw [← normalize_normalized_factor _ hp, ← normalize_normalized_factor _ hq] <;> - exact - normalize_eq_normalize hdvd - ((prime_of_normalized_factor _ hp).irreducible.dvd_symm - (prime_of_normalized_factor _ hq).irreducible hdvd) - -end Multiplicative - -end UniqueFactorizationMonoid - -namespace Associates - -open UniqueFactorizationMonoid Associated Multiset - -variable [CancelCommMonoidWithZero α] - -/-- `FactorSet α` representation elements of unique factorization domain as multisets. -`Multiset α` produced by `normalizedFactors` are only unique up to associated elements, while the -multisets in `FactorSet α` are unique by equality and restricted to irreducible elements. This -gives us a representation of each element as a unique multisets (or the added ⊤ for 0), which has a -complete lattice structure. Infimum is the greatest common divisor and supremum is the least common -multiple. --/ -abbrev FactorSet.{u} (α : Type u) [CancelCommMonoidWithZero α] : Type u := - WithTop (Multiset { a : Associates α // Irreducible a }) - -attribute [local instance] Associated.setoid - -theorem FactorSet.coe_add {a b : Multiset { a : Associates α // Irreducible a }} : - (↑(a + b) : FactorSet α) = a + b := by norm_cast - -theorem FactorSet.sup_add_inf_eq_add [DecidableEq (Associates α)] : - ∀ a b : FactorSet α, a ⊔ b + a ⊓ b = a + b - | ⊤, b => show ⊤ ⊔ b + ⊤ ⊓ b = ⊤ + b by simp - | a, ⊤ => show a ⊔ ⊤ + a ⊓ ⊤ = a + ⊤ by simp - | WithTop.some a, WithTop.some b => - show (a : FactorSet α) ⊔ b + (a : FactorSet α) ⊓ b = a + b by - rw [← WithTop.coe_sup, ← WithTop.coe_inf, ← WithTop.coe_add, ← WithTop.coe_add, - WithTop.coe_eq_coe] - exact Multiset.union_add_inter _ _ - -/-- Evaluates the product of a `FactorSet` to be the product of the corresponding multiset, - or `0` if there is none. -/ -def FactorSet.prod : FactorSet α → Associates α - | ⊤ => 0 - | WithTop.some s => (s.map (↑)).prod - -@[simp] -theorem prod_top : (⊤ : FactorSet α).prod = 0 := - rfl - -@[simp] -theorem prod_coe {s : Multiset { a : Associates α // Irreducible a }} : - FactorSet.prod (s : FactorSet α) = (s.map (↑)).prod := - rfl - -@[simp] -theorem prod_add : ∀ a b : FactorSet α, (a + b).prod = a.prod * b.prod - | ⊤, b => show (⊤ + b).prod = (⊤ : FactorSet α).prod * b.prod by simp - | a, ⊤ => show (a + ⊤).prod = a.prod * (⊤ : FactorSet α).prod by simp - | WithTop.some a, WithTop.some b => by - rw [← FactorSet.coe_add, prod_coe, prod_coe, prod_coe, Multiset.map_add, Multiset.prod_add] - -@[gcongr] -theorem prod_mono : ∀ {a b : FactorSet α}, a ≤ b → a.prod ≤ b.prod - | ⊤, b, h => by - have : b = ⊤ := top_unique h - rw [this, prod_top] - | a, ⊤, _ => show a.prod ≤ (⊤ : FactorSet α).prod by simp - | WithTop.some _, WithTop.some _, h => - prod_le_prod <| Multiset.map_le_map <| WithTop.coe_le_coe.1 <| h - -theorem FactorSet.prod_eq_zero_iff [Nontrivial α] (p : FactorSet α) : p.prod = 0 ↔ p = ⊤ := by - unfold FactorSet at p - induction p -- TODO: `induction_eliminator` doesn't work with `abbrev` - · simp only [eq_self_iff_true, Associates.prod_top] - · rw [prod_coe, Multiset.prod_eq_zero_iff, Multiset.mem_map, eq_false WithTop.coe_ne_top, - iff_false, not_exists] - exact fun a => not_and_of_not_right _ a.prop.ne_zero - -section count - -variable [DecidableEq (Associates α)] - -/-- `bcount p s` is the multiplicity of `p` in the FactorSet `s` (with bundled `p`)-/ -def bcount (p : { a : Associates α // Irreducible a }) : - FactorSet α → ℕ - | ⊤ => 0 - | WithTop.some s => s.count p - -variable [∀ p : Associates α, Decidable (Irreducible p)] {p : Associates α} - -/-- `count p s` is the multiplicity of the irreducible `p` in the FactorSet `s`. - -If `p` is not irreducible, `count p s` is defined to be `0`. -/ -def count (p : Associates α) : FactorSet α → ℕ := - if hp : Irreducible p then bcount ⟨p, hp⟩ else 0 - -@[simp] -theorem count_some (hp : Irreducible p) (s : Multiset _) : - count p (WithTop.some s) = s.count ⟨p, hp⟩ := by - simp only [count, dif_pos hp, bcount] - -@[simp] -theorem count_zero (hp : Irreducible p) : count p (0 : FactorSet α) = 0 := by - simp only [count, dif_pos hp, bcount, Multiset.count_zero] - -theorem count_reducible (hp : ¬Irreducible p) : count p = 0 := dif_neg hp - -end count - -section Mem - -/-- membership in a FactorSet (bundled version) -/ -def BfactorSetMem : { a : Associates α // Irreducible a } → FactorSet α → Prop - | _, ⊤ => True - | p, some l => p ∈ l - -/-- `FactorSetMem p s` is the predicate that the irreducible `p` is a member of -`s : FactorSet α`. - -If `p` is not irreducible, `p` is not a member of any `FactorSet`. -/ -def FactorSetMem (s : FactorSet α) (p : Associates α) : Prop := - letI : Decidable (Irreducible p) := Classical.dec _ - if hp : Irreducible p then BfactorSetMem ⟨p, hp⟩ s else False - -instance : Membership (Associates α) (FactorSet α) := - ⟨FactorSetMem⟩ - -@[simp] -theorem factorSetMem_eq_mem (p : Associates α) (s : FactorSet α) : FactorSetMem s p = (p ∈ s) := - rfl - -theorem mem_factorSet_top {p : Associates α} {hp : Irreducible p} : p ∈ (⊤ : FactorSet α) := by - dsimp only [Membership.mem]; dsimp only [FactorSetMem]; split_ifs; exact trivial - -theorem mem_factorSet_some {p : Associates α} {hp : Irreducible p} - {l : Multiset { a : Associates α // Irreducible a }} : - p ∈ (l : FactorSet α) ↔ Subtype.mk p hp ∈ l := by - dsimp only [Membership.mem]; dsimp only [FactorSetMem]; split_ifs; rfl - -theorem reducible_not_mem_factorSet {p : Associates α} (hp : ¬Irreducible p) (s : FactorSet α) : - ¬p ∈ s := fun h ↦ by - rwa [← factorSetMem_eq_mem, FactorSetMem, dif_neg hp] at h - -theorem irreducible_of_mem_factorSet {p : Associates α} {s : FactorSet α} (h : p ∈ s) : - Irreducible p := - by_contra fun hp ↦ reducible_not_mem_factorSet hp s h - -end Mem - -variable [UniqueFactorizationMonoid α] - -theorem unique' {p q : Multiset (Associates α)} : - (∀ a ∈ p, Irreducible a) → (∀ a ∈ q, Irreducible a) → p.prod = q.prod → p = q := by - apply Multiset.induction_on_multiset_quot p - apply Multiset.induction_on_multiset_quot q - intro s t hs ht eq - refine Multiset.map_mk_eq_map_mk_of_rel (UniqueFactorizationMonoid.factors_unique ?_ ?_ ?_) - · exact fun a ha => irreducible_mk.1 <| hs _ <| Multiset.mem_map_of_mem _ ha - · exact fun a ha => irreducible_mk.1 <| ht _ <| Multiset.mem_map_of_mem _ ha - have eq' : (Quot.mk Setoid.r : α → Associates α) = Associates.mk := funext quot_mk_eq_mk - rwa [eq', prod_mk, prod_mk, mk_eq_mk_iff_associated] at eq - -theorem FactorSet.unique [Nontrivial α] {p q : FactorSet α} (h : p.prod = q.prod) : p = q := by - -- TODO: `induction_eliminator` doesn't work with `abbrev` - unfold FactorSet at p q - induction p <;> induction q - · rfl - · rw [eq_comm, ← FactorSet.prod_eq_zero_iff, ← h, Associates.prod_top] - · rw [← FactorSet.prod_eq_zero_iff, h, Associates.prod_top] - · congr 1 - rw [← Multiset.map_eq_map Subtype.coe_injective] - apply unique' _ _ h <;> - · intro a ha - obtain ⟨⟨a', irred⟩, -, rfl⟩ := Multiset.mem_map.mp ha - rwa [Subtype.coe_mk] - -theorem prod_le_prod_iff_le [Nontrivial α] {p q : Multiset (Associates α)} - (hp : ∀ a ∈ p, Irreducible a) (hq : ∀ a ∈ q, Irreducible a) : p.prod ≤ q.prod ↔ p ≤ q := by - refine ⟨?_, prod_le_prod⟩ - rintro ⟨c, eqc⟩ - refine Multiset.le_iff_exists_add.2 ⟨factors c, unique' hq (fun x hx ↦ ?_) ?_⟩ - · obtain h | h := Multiset.mem_add.1 hx - · exact hp x h - · exact irreducible_of_factor _ h - · rw [eqc, Multiset.prod_add] - congr - refine associated_iff_eq.mp (factors_prod fun hc => ?_).symm - refine not_irreducible_zero (hq _ ?_) - rw [← prod_eq_zero_iff, eqc, hc, mul_zero] - -/-- This returns the multiset of irreducible factors as a `FactorSet`, - a multiset of irreducible associates `WithTop`. -/ -noncomputable def factors' (a : α) : Multiset { a : Associates α // Irreducible a } := - (factors a).pmap (fun a ha => ⟨Associates.mk a, irreducible_mk.2 ha⟩) irreducible_of_factor - -@[simp] -theorem map_subtype_coe_factors' {a : α} : - (factors' a).map (↑) = (factors a).map Associates.mk := by - simp [factors', Multiset.map_pmap, Multiset.pmap_eq_map] - -theorem factors'_cong {a b : α} (h : a ~ᵤ b) : factors' a = factors' b := by - obtain rfl | hb := eq_or_ne b 0 - · rw [associated_zero_iff_eq_zero] at h - rw [h] - have ha : a ≠ 0 := by - contrapose! hb with ha - rw [← associated_zero_iff_eq_zero, ← ha] - exact h.symm - rw [← Multiset.map_eq_map Subtype.coe_injective, map_subtype_coe_factors', - map_subtype_coe_factors', ← rel_associated_iff_map_eq_map] - exact - factors_unique irreducible_of_factor irreducible_of_factor - ((factors_prod ha).trans <| h.trans <| (factors_prod hb).symm) - -/-- This returns the multiset of irreducible factors of an associate as a `FactorSet`, - a multiset of irreducible associates `WithTop`. -/ -noncomputable def factors (a : Associates α) : FactorSet α := by - classical refine if h : a = 0 then ⊤ else Quotient.hrecOn a (fun x _ => factors' x) ?_ h - intro a b hab - apply Function.hfunext - · have : a ~ᵤ 0 ↔ b ~ᵤ 0 := Iff.intro (fun ha0 => hab.symm.trans ha0) fun hb0 => hab.trans hb0 - simp only [associated_zero_iff_eq_zero] at this - simp only [quotient_mk_eq_mk, this, mk_eq_zero] - exact fun ha hb _ => heq_of_eq <| congr_arg some <| factors'_cong hab - -@[simp] -theorem factors_zero : (0 : Associates α).factors = ⊤ := - dif_pos rfl - -@[deprecated (since := "2024-03-16")] alias factors_0 := factors_zero - -@[simp] -theorem factors_mk (a : α) (h : a ≠ 0) : (Associates.mk a).factors = factors' a := by - classical - apply dif_neg - apply mt mk_eq_zero.1 h - -@[simp] -theorem factors_prod (a : Associates α) : a.factors.prod = a := by - rcases Associates.mk_surjective a with ⟨a, rfl⟩ - rcases eq_or_ne a 0 with rfl | ha - · simp - · simp [ha, prod_mk, mk_eq_mk_iff_associated, UniqueFactorizationMonoid.factors_prod, - -Quotient.eq] - -@[simp] -theorem prod_factors [Nontrivial α] (s : FactorSet α) : s.prod.factors = s := - FactorSet.unique <| factors_prod _ - -@[nontriviality] -theorem factors_subsingleton [Subsingleton α] {a : Associates α} : a.factors = ⊤ := by - have : Subsingleton (Associates α) := inferInstance - convert factors_zero - -theorem factors_eq_top_iff_zero {a : Associates α} : a.factors = ⊤ ↔ a = 0 := by - nontriviality α - exact ⟨fun h ↦ by rwa [← factors_prod a, FactorSet.prod_eq_zero_iff], fun h ↦ h ▸ factors_zero⟩ - -@[deprecated (since := "2024-04-16")] alias factors_eq_none_iff_zero := factors_eq_top_iff_zero - -theorem factors_eq_some_iff_ne_zero {a : Associates α} : - (∃ s : Multiset { p : Associates α // Irreducible p }, a.factors = s) ↔ a ≠ 0 := by - simp_rw [@eq_comm _ a.factors, ← WithTop.ne_top_iff_exists] - exact factors_eq_top_iff_zero.not - -theorem eq_of_factors_eq_factors {a b : Associates α} (h : a.factors = b.factors) : a = b := by - have : a.factors.prod = b.factors.prod := by rw [h] - rwa [factors_prod, factors_prod] at this - -theorem eq_of_prod_eq_prod [Nontrivial α] {a b : FactorSet α} (h : a.prod = b.prod) : a = b := by - have : a.prod.factors = b.prod.factors := by rw [h] - rwa [prod_factors, prod_factors] at this - -@[simp] -theorem factors_mul (a b : Associates α) : (a * b).factors = a.factors + b.factors := by - nontriviality α - refine eq_of_prod_eq_prod <| eq_of_factors_eq_factors ?_ - rw [prod_add, factors_prod, factors_prod, factors_prod] - -@[gcongr] -theorem factors_mono : ∀ {a b : Associates α}, a ≤ b → a.factors ≤ b.factors - | s, t, ⟨d, eq⟩ => by rw [eq, factors_mul]; exact le_add_of_nonneg_right bot_le - -@[simp] -theorem factors_le {a b : Associates α} : a.factors ≤ b.factors ↔ a ≤ b := by - refine ⟨fun h ↦ ?_, factors_mono⟩ - have : a.factors.prod ≤ b.factors.prod := prod_mono h - rwa [factors_prod, factors_prod] at this - -section count - -variable [DecidableEq (Associates α)] [∀ p : Associates α, Decidable (Irreducible p)] - -theorem eq_factors_of_eq_counts {a b : Associates α} (ha : a ≠ 0) (hb : b ≠ 0) - (h : ∀ p : Associates α, Irreducible p → p.count a.factors = p.count b.factors) : - a.factors = b.factors := by - obtain ⟨sa, h_sa⟩ := factors_eq_some_iff_ne_zero.mpr ha - obtain ⟨sb, h_sb⟩ := factors_eq_some_iff_ne_zero.mpr hb - rw [h_sa, h_sb] at h ⊢ - rw [WithTop.coe_eq_coe] - have h_count : ∀ (p : Associates α) (hp : Irreducible p), - sa.count ⟨p, hp⟩ = sb.count ⟨p, hp⟩ := by - intro p hp - rw [← count_some, ← count_some, h p hp] - apply Multiset.toFinsupp.injective - ext ⟨p, hp⟩ - rw [Multiset.toFinsupp_apply, Multiset.toFinsupp_apply, h_count p hp] - -theorem eq_of_eq_counts {a b : Associates α} (ha : a ≠ 0) (hb : b ≠ 0) - (h : ∀ p : Associates α, Irreducible p → p.count a.factors = p.count b.factors) : a = b := - eq_of_factors_eq_factors (eq_factors_of_eq_counts ha hb h) - -theorem count_le_count_of_factors_le {a b p : Associates α} (hb : b ≠ 0) (hp : Irreducible p) - (h : a.factors ≤ b.factors) : p.count a.factors ≤ p.count b.factors := by - by_cases ha : a = 0 - · simp_all - obtain ⟨sa, h_sa⟩ := factors_eq_some_iff_ne_zero.mpr ha - obtain ⟨sb, h_sb⟩ := factors_eq_some_iff_ne_zero.mpr hb - rw [h_sa, h_sb] at h ⊢ - rw [count_some hp, count_some hp]; rw [WithTop.coe_le_coe] at h - exact Multiset.count_le_of_le _ h - -theorem count_le_count_of_le {a b p : Associates α} (hb : b ≠ 0) (hp : Irreducible p) (h : a ≤ b) : - p.count a.factors ≤ p.count b.factors := - count_le_count_of_factors_le hb hp <| factors_mono h - -end count - -theorem prod_le [Nontrivial α] {a b : FactorSet α} : a.prod ≤ b.prod ↔ a ≤ b := by - refine ⟨fun h ↦ ?_, prod_mono⟩ - have : a.prod.factors ≤ b.prod.factors := factors_mono h - rwa [prod_factors, prod_factors] at this - -open Classical in -noncomputable instance : Max (Associates α) := - ⟨fun a b => (a.factors ⊔ b.factors).prod⟩ - -open Classical in -noncomputable instance : Min (Associates α) := - ⟨fun a b => (a.factors ⊓ b.factors).prod⟩ - -open Classical in -noncomputable instance : Lattice (Associates α) := - { Associates.instPartialOrder with - sup := (· ⊔ ·) - inf := (· ⊓ ·) - sup_le := fun _ _ c hac hbc => - factors_prod c ▸ prod_mono (sup_le (factors_mono hac) (factors_mono hbc)) - le_sup_left := fun a _ => le_trans (le_of_eq (factors_prod a).symm) <| prod_mono <| le_sup_left - le_sup_right := fun _ b => - le_trans (le_of_eq (factors_prod b).symm) <| prod_mono <| le_sup_right - le_inf := fun a _ _ hac hbc => - factors_prod a ▸ prod_mono (le_inf (factors_mono hac) (factors_mono hbc)) - inf_le_left := fun a _ => le_trans (prod_mono inf_le_left) (le_of_eq (factors_prod a)) - inf_le_right := fun _ b => le_trans (prod_mono inf_le_right) (le_of_eq (factors_prod b)) } - -open Classical in -theorem sup_mul_inf (a b : Associates α) : (a ⊔ b) * (a ⊓ b) = a * b := - show (a.factors ⊔ b.factors).prod * (a.factors ⊓ b.factors).prod = a * b by - nontriviality α - refine eq_of_factors_eq_factors ?_ - rw [← prod_add, prod_factors, factors_mul, FactorSet.sup_add_inf_eq_add] - -theorem dvd_of_mem_factors {a p : Associates α} (hm : p ∈ factors a) : - p ∣ a := by - rcases eq_or_ne a 0 with rfl | ha0 - · exact dvd_zero p - obtain ⟨a0, nza, ha'⟩ := exists_non_zero_rep ha0 - rw [← Associates.factors_prod a] - rw [← ha', factors_mk a0 nza] at hm ⊢ - rw [prod_coe] - apply Multiset.dvd_prod; apply Multiset.mem_map.mpr - exact ⟨⟨p, irreducible_of_mem_factorSet hm⟩, mem_factorSet_some.mp hm, rfl⟩ - -theorem dvd_of_mem_factors' {a : α} {p : Associates α} {hp : Irreducible p} {hz : a ≠ 0} - (h_mem : Subtype.mk p hp ∈ factors' a) : p ∣ Associates.mk a := by - haveI := Classical.decEq (Associates α) - apply dvd_of_mem_factors - rw [factors_mk _ hz] - apply mem_factorSet_some.2 h_mem - -theorem mem_factors'_of_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) (hd : p ∣ a) : - Subtype.mk (Associates.mk p) (irreducible_mk.2 hp) ∈ factors' a := by - obtain ⟨q, hq, hpq⟩ := exists_mem_factors_of_dvd ha0 hp hd - apply Multiset.mem_pmap.mpr; use q; use hq - exact Subtype.eq (Eq.symm (mk_eq_mk_iff_associated.mpr hpq)) - -theorem mem_factors'_iff_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) : - Subtype.mk (Associates.mk p) (irreducible_mk.2 hp) ∈ factors' a ↔ p ∣ a := by - constructor - · rw [← mk_dvd_mk] - apply dvd_of_mem_factors' - apply ha0 - · apply mem_factors'_of_dvd ha0 hp - -theorem mem_factors_of_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) (hd : p ∣ a) : - Associates.mk p ∈ factors (Associates.mk a) := by - rw [factors_mk _ ha0] - exact mem_factorSet_some.mpr (mem_factors'_of_dvd ha0 hp hd) - -theorem mem_factors_iff_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) : - Associates.mk p ∈ factors (Associates.mk a) ↔ p ∣ a := by - constructor - · rw [← mk_dvd_mk] - apply dvd_of_mem_factors - · apply mem_factors_of_dvd ha0 hp - -open Classical in -theorem exists_prime_dvd_of_not_inf_one {a b : α} (ha : a ≠ 0) (hb : b ≠ 0) - (h : Associates.mk a ⊓ Associates.mk b ≠ 1) : ∃ p : α, Prime p ∧ p ∣ a ∧ p ∣ b := by - have hz : factors (Associates.mk a) ⊓ factors (Associates.mk b) ≠ 0 := by - contrapose! h with hf - change (factors (Associates.mk a) ⊓ factors (Associates.mk b)).prod = 1 - rw [hf] - exact Multiset.prod_zero - rw [factors_mk a ha, factors_mk b hb, ← WithTop.coe_inf] at hz - obtain ⟨⟨p0, p0_irr⟩, p0_mem⟩ := Multiset.exists_mem_of_ne_zero ((mt WithTop.coe_eq_coe.mpr) hz) - rw [Multiset.inf_eq_inter] at p0_mem - obtain ⟨p, rfl⟩ : ∃ p, Associates.mk p = p0 := Quot.exists_rep p0 - refine ⟨p, ?_, ?_, ?_⟩ - · rw [← UniqueFactorizationMonoid.irreducible_iff_prime, ← irreducible_mk] - exact p0_irr - · apply dvd_of_mk_le_mk - apply dvd_of_mem_factors' (Multiset.mem_inter.mp p0_mem).left - apply ha - · apply dvd_of_mk_le_mk - apply dvd_of_mem_factors' (Multiset.mem_inter.mp p0_mem).right - apply hb - -theorem coprime_iff_inf_one {a b : α} (ha0 : a ≠ 0) (hb0 : b ≠ 0) : - Associates.mk a ⊓ Associates.mk b = 1 ↔ ∀ {d : α}, d ∣ a → d ∣ b → ¬Prime d := by - constructor - · intro hg p ha hb hp - refine (Associates.prime_mk.mpr hp).not_unit (isUnit_of_dvd_one ?_) - rw [← hg] - exact le_inf (mk_le_mk_of_dvd ha) (mk_le_mk_of_dvd hb) - · contrapose - intro hg hc - obtain ⟨p, hp, hpa, hpb⟩ := exists_prime_dvd_of_not_inf_one ha0 hb0 hg - exact hc hpa hpb hp - -theorem factors_self [Nontrivial α] {p : Associates α} (hp : Irreducible p) : - p.factors = WithTop.some {⟨p, hp⟩} := - eq_of_prod_eq_prod - (by rw [factors_prod, FactorSet.prod.eq_def]; dsimp; rw [prod_singleton]) - -theorem factors_prime_pow [Nontrivial α] {p : Associates α} (hp : Irreducible p) (k : ℕ) : - factors (p ^ k) = WithTop.some (Multiset.replicate k ⟨p, hp⟩) := - eq_of_prod_eq_prod - (by - rw [Associates.factors_prod, FactorSet.prod.eq_def] - dsimp; rw [Multiset.map_replicate, Multiset.prod_replicate, Subtype.coe_mk]) - -theorem prime_pow_le_iff_le_bcount [DecidableEq (Associates α)] {m p : Associates α} - (h₁ : m ≠ 0) (h₂ : Irreducible p) {k : ℕ} : p ^ k ≤ m ↔ k ≤ bcount ⟨p, h₂⟩ m.factors := by - rcases Associates.exists_non_zero_rep h₁ with ⟨m, hm, rfl⟩ - have := nontrivial_of_ne _ _ hm - rw [bcount.eq_def, factors_mk, Multiset.le_count_iff_replicate_le, ← factors_le, - factors_prime_pow, factors_mk, WithTop.coe_le_coe] <;> assumption - -@[simp] -theorem factors_one [Nontrivial α] : factors (1 : Associates α) = 0 := by - apply eq_of_prod_eq_prod - rw [Associates.factors_prod] - exact Multiset.prod_zero - -@[simp] -theorem pow_factors [Nontrivial α] {a : Associates α} {k : ℕ} : - (a ^ k).factors = k • a.factors := by - induction' k with n h - · rw [zero_nsmul, pow_zero] - exact factors_one - · rw [pow_succ, succ_nsmul, factors_mul, h] - -section count - -variable [DecidableEq (Associates α)] [∀ p : Associates α, Decidable (Irreducible p)] - -theorem prime_pow_dvd_iff_le {m p : Associates α} (h₁ : m ≠ 0) (h₂ : Irreducible p) {k : ℕ} : - p ^ k ≤ m ↔ k ≤ count p m.factors := by - rw [count, dif_pos h₂, prime_pow_le_iff_le_bcount h₁] - -theorem le_of_count_ne_zero {m p : Associates α} (h0 : m ≠ 0) (hp : Irreducible p) : - count p m.factors ≠ 0 → p ≤ m := by - nontriviality α - rw [← pos_iff_ne_zero] - intro h - rw [← pow_one p] - apply (prime_pow_dvd_iff_le h0 hp).2 - simpa only - -theorem count_ne_zero_iff_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) : - (Associates.mk p).count (Associates.mk a).factors ≠ 0 ↔ p ∣ a := by - nontriviality α - rw [← Associates.mk_le_mk_iff_dvd] - refine - ⟨fun h => - Associates.le_of_count_ne_zero (Associates.mk_ne_zero.mpr ha0) - (Associates.irreducible_mk.mpr hp) h, - fun h => ?_⟩ - rw [← pow_one (Associates.mk p), - Associates.prime_pow_dvd_iff_le (Associates.mk_ne_zero.mpr ha0) - (Associates.irreducible_mk.mpr hp)] at h - exact (zero_lt_one.trans_le h).ne' - -theorem count_self [Nontrivial α] {p : Associates α} - (hp : Irreducible p) : p.count p.factors = 1 := by - simp [factors_self hp, Associates.count_some hp] - -theorem count_eq_zero_of_ne {p q : Associates α} (hp : Irreducible p) - (hq : Irreducible q) (h : p ≠ q) : p.count q.factors = 0 := - not_ne_iff.mp fun h' ↦ h <| associated_iff_eq.mp <| hp.associated_of_dvd hq <| - le_of_count_ne_zero hq.ne_zero hp h' - -theorem count_mul {a : Associates α} (ha : a ≠ 0) {b : Associates α} - (hb : b ≠ 0) {p : Associates α} (hp : Irreducible p) : - count p (factors (a * b)) = count p a.factors + count p b.factors := by - obtain ⟨a0, nza, rfl⟩ := exists_non_zero_rep ha - obtain ⟨b0, nzb, rfl⟩ := exists_non_zero_rep hb - rw [factors_mul, factors_mk a0 nza, factors_mk b0 nzb, ← FactorSet.coe_add, count_some hp, - Multiset.count_add, count_some hp, count_some hp] - -theorem count_of_coprime {a : Associates α} (ha : a ≠ 0) - {b : Associates α} (hb : b ≠ 0) (hab : ∀ d, d ∣ a → d ∣ b → ¬Prime d) {p : Associates α} - (hp : Irreducible p) : count p a.factors = 0 ∨ count p b.factors = 0 := by - rw [or_iff_not_imp_left, ← Ne] - intro hca - contrapose! hab with hcb - exact ⟨p, le_of_count_ne_zero ha hp hca, le_of_count_ne_zero hb hp hcb, - UniqueFactorizationMonoid.irreducible_iff_prime.mp hp⟩ - -theorem count_mul_of_coprime {a : Associates α} {b : Associates α} - (hb : b ≠ 0) {p : Associates α} (hp : Irreducible p) (hab : ∀ d, d ∣ a → d ∣ b → ¬Prime d) : - count p a.factors = 0 ∨ count p a.factors = count p (a * b).factors := by - by_cases ha : a = 0 - · simp [ha] - cases' count_of_coprime ha hb hab hp with hz hb0; · tauto - apply Or.intro_right - rw [count_mul ha hb hp, hb0, add_zero] - -theorem count_mul_of_coprime' {a b : Associates α} {p : Associates α} - (hp : Irreducible p) (hab : ∀ d, d ∣ a → d ∣ b → ¬Prime d) : - count p (a * b).factors = count p a.factors ∨ count p (a * b).factors = count p b.factors := by - by_cases ha : a = 0 - · simp [ha] - by_cases hb : b = 0 - · simp [hb] - rw [count_mul ha hb hp] - cases' count_of_coprime ha hb hab hp with ha0 hb0 - · apply Or.intro_right - rw [ha0, zero_add] - · apply Or.intro_left - rw [hb0, add_zero] - -theorem dvd_count_of_dvd_count_mul {a b : Associates α} (hb : b ≠ 0) - {p : Associates α} (hp : Irreducible p) (hab : ∀ d, d ∣ a → d ∣ b → ¬Prime d) {k : ℕ} - (habk : k ∣ count p (a * b).factors) : k ∣ count p a.factors := by - by_cases ha : a = 0 - · simpa [*] using habk - cases' count_of_coprime ha hb hab hp with hz h - · rw [hz] - exact dvd_zero k - · rw [count_mul ha hb hp, h] at habk - exact habk - -theorem count_pow [Nontrivial α] {a : Associates α} (ha : a ≠ 0) - {p : Associates α} (hp : Irreducible p) (k : ℕ) : - count p (a ^ k).factors = k * count p a.factors := by - induction' k with n h - · rw [pow_zero, factors_one, zero_mul, count_zero hp] - · rw [pow_succ', count_mul ha (pow_ne_zero _ ha) hp, h] - ring - -theorem dvd_count_pow [Nontrivial α] {a : Associates α} (ha : a ≠ 0) - {p : Associates α} (hp : Irreducible p) (k : ℕ) : k ∣ count p (a ^ k).factors := by - rw [count_pow ha hp] - apply dvd_mul_right - -theorem is_pow_of_dvd_count {a : Associates α} - (ha : a ≠ 0) {k : ℕ} (hk : ∀ p : Associates α, Irreducible p → k ∣ count p a.factors) : - ∃ b : Associates α, a = b ^ k := by - nontriviality α - obtain ⟨a0, hz, rfl⟩ := exists_non_zero_rep ha - rw [factors_mk a0 hz] at hk - have hk' : ∀ p, p ∈ factors' a0 → k ∣ (factors' a0).count p := by - rintro p - - have pp : p = ⟨p.val, p.2⟩ := by simp only [Subtype.coe_eta] - rw [pp, ← count_some p.2] - exact hk p.val p.2 - obtain ⟨u, hu⟩ := Multiset.exists_smul_of_dvd_count _ hk' - use FactorSet.prod (u : FactorSet α) - apply eq_of_factors_eq_factors - rw [pow_factors, prod_factors, factors_mk a0 hz, hu] - exact WithBot.coe_nsmul u k - -/-- The only divisors of prime powers are prime powers. See `eq_pow_find_of_dvd_irreducible_pow` -for an explicit expression as a p-power (without using `count`). -/ -theorem eq_pow_count_factors_of_dvd_pow {p a : Associates α} - (hp : Irreducible p) {n : ℕ} (h : a ∣ p ^ n) : a = p ^ p.count a.factors := by - nontriviality α - have hph := pow_ne_zero n hp.ne_zero - have ha := ne_zero_of_dvd_ne_zero hph h - apply eq_of_eq_counts ha (pow_ne_zero _ hp.ne_zero) - have eq_zero_of_ne : ∀ q : Associates α, Irreducible q → q ≠ p → _ = 0 := fun q hq h' => - Nat.eq_zero_of_le_zero <| by - convert count_le_count_of_le hph hq h - symm - rw [count_pow hp.ne_zero hq, count_eq_zero_of_ne hq hp h', mul_zero] - intro q hq - rw [count_pow hp.ne_zero hq] - by_cases h : q = p - · rw [h, count_self hp, mul_one] - · rw [count_eq_zero_of_ne hq hp h, mul_zero, eq_zero_of_ne q hq h] - -theorem count_factors_eq_find_of_dvd_pow {a p : Associates α} - (hp : Irreducible p) [∀ n : ℕ, Decidable (a ∣ p ^ n)] {n : ℕ} (h : a ∣ p ^ n) : - @Nat.find (fun n => a ∣ p ^ n) _ ⟨n, h⟩ = p.count a.factors := by - apply le_antisymm - · refine Nat.find_le ⟨1, ?_⟩ - rw [mul_one] - symm - exact eq_pow_count_factors_of_dvd_pow hp h - · have hph := pow_ne_zero (@Nat.find (fun n => a ∣ p ^ n) _ ⟨n, h⟩) hp.ne_zero - cases' subsingleton_or_nontrivial α with hα hα - · simp [eq_iff_true_of_subsingleton] at hph - convert count_le_count_of_le hph hp (@Nat.find_spec (fun n => a ∣ p ^ n) _ ⟨n, h⟩) - rw [count_pow hp.ne_zero hp, count_self hp, mul_one] - -end count - -theorem eq_pow_of_mul_eq_pow {a b c : Associates α} (ha : a ≠ 0) (hb : b ≠ 0) - (hab : ∀ d, d ∣ a → d ∣ b → ¬Prime d) {k : ℕ} (h : a * b = c ^ k) : - ∃ d : Associates α, a = d ^ k := by - classical - nontriviality α - by_cases hk0 : k = 0 - · use 1 - rw [hk0, pow_zero] at h ⊢ - apply (mul_eq_one.1 h).1 - · refine is_pow_of_dvd_count ha fun p hp ↦ ?_ - apply dvd_count_of_dvd_count_mul hb hp hab - rw [h] - apply dvd_count_pow _ hp - rintro rfl - rw [zero_pow hk0] at h - cases mul_eq_zero.mp h <;> contradiction - -/-- The only divisors of prime powers are prime powers. -/ -theorem eq_pow_find_of_dvd_irreducible_pow {a p : Associates α} (hp : Irreducible p) - [∀ n : ℕ, Decidable (a ∣ p ^ n)] {n : ℕ} (h : a ∣ p ^ n) : - a = p ^ @Nat.find (fun n => a ∣ p ^ n) _ ⟨n, h⟩ := by - classical rw [count_factors_eq_find_of_dvd_pow hp, ← eq_pow_count_factors_of_dvd_pow hp h] - exact h - -end Associates - -section - -open Associates UniqueFactorizationMonoid - -/-- `toGCDMonoid` constructs a GCD monoid out of a unique factorization domain. -/ -noncomputable def UniqueFactorizationMonoid.toGCDMonoid (α : Type*) [CancelCommMonoidWithZero α] - [UniqueFactorizationMonoid α] : GCDMonoid α where - gcd a b := Quot.out (Associates.mk a ⊓ Associates.mk b : Associates α) - lcm a b := Quot.out (Associates.mk a ⊔ Associates.mk b : Associates α) - gcd_dvd_left a b := by - rw [← mk_dvd_mk, Associates.quot_out, congr_fun₂ dvd_eq_le] - exact inf_le_left - gcd_dvd_right a b := by - rw [← mk_dvd_mk, Associates.quot_out, congr_fun₂ dvd_eq_le] - exact inf_le_right - dvd_gcd {a b c} hac hab := by - rw [← mk_dvd_mk, Associates.quot_out, congr_fun₂ dvd_eq_le, le_inf_iff, - mk_le_mk_iff_dvd, mk_le_mk_iff_dvd] - exact ⟨hac, hab⟩ - lcm_zero_left a := by simp - lcm_zero_right a := by simp - gcd_mul_lcm a b := by - rw [← mk_eq_mk_iff_associated, ← Associates.mk_mul_mk, ← associated_iff_eq, Associates.quot_out, - Associates.quot_out, mul_comm, sup_mul_inf, Associates.mk_mul_mk] - -/-- `toNormalizedGCDMonoid` constructs a GCD monoid out of a normalization on a - unique factorization domain. -/ -noncomputable def UniqueFactorizationMonoid.toNormalizedGCDMonoid (α : Type*) - [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] [NormalizationMonoid α] : - NormalizedGCDMonoid α := - { ‹NormalizationMonoid α› with - gcd := fun a b => (Associates.mk a ⊓ Associates.mk b).out - lcm := fun a b => (Associates.mk a ⊔ Associates.mk b).out - gcd_dvd_left := fun a b => (out_dvd_iff a (Associates.mk a ⊓ Associates.mk b)).2 <| inf_le_left - gcd_dvd_right := fun a b => - (out_dvd_iff b (Associates.mk a ⊓ Associates.mk b)).2 <| inf_le_right - dvd_gcd := fun {a} {b} {c} hac hab => - show a ∣ (Associates.mk c ⊓ Associates.mk b).out by - rw [dvd_out_iff, le_inf_iff, mk_le_mk_iff_dvd, mk_le_mk_iff_dvd] - exact ⟨hac, hab⟩ - lcm_zero_left := fun a => show (⊤ ⊔ Associates.mk a).out = 0 by simp - lcm_zero_right := fun a => show (Associates.mk a ⊔ ⊤).out = 0 by simp - gcd_mul_lcm := fun a b => by - rw [← out_mul, mul_comm, sup_mul_inf, mk_mul_mk, out_mk] - exact normalize_associated (a * b) - normalize_gcd := fun a b => by apply normalize_out _ - normalize_lcm := fun a b => by apply normalize_out _ } - -instance (α) [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] : - Nonempty (NormalizedGCDMonoid α) := by - letI := UniqueFactorizationMonoid.normalizationMonoid (α := α) - classical exact ⟨UniqueFactorizationMonoid.toNormalizedGCDMonoid α⟩ - -end - -namespace UniqueFactorizationMonoid - -/-- If `y` is a nonzero element of a unique factorization monoid with finitely -many units (e.g. `ℤ`, `Ideal (ring_of_integers K)`), it has finitely many divisors. -/ -noncomputable def fintypeSubtypeDvd {M : Type*} [CancelCommMonoidWithZero M] - [UniqueFactorizationMonoid M] [Fintype Mˣ] (y : M) (hy : y ≠ 0) : Fintype { x // x ∣ y } := by - haveI : Nontrivial M := ⟨⟨y, 0, hy⟩⟩ - haveI : NormalizationMonoid M := UniqueFactorizationMonoid.normalizationMonoid - haveI := Classical.decEq M - haveI := Classical.decEq (Associates M) - -- We'll show `fun (u : Mˣ) (f ⊆ factors y) ↦ u * Π f` is injective - -- and has image exactly the divisors of `y`. - refine - Fintype.ofFinset - (((normalizedFactors y).powerset.toFinset ×ˢ (Finset.univ : Finset Mˣ)).image fun s => - (s.snd : M) * s.fst.prod) - fun x => ?_ - simp only [exists_prop, Finset.mem_image, Finset.mem_product, Finset.mem_univ, and_true, - Multiset.mem_toFinset, Multiset.mem_powerset, exists_eq_right, Multiset.mem_map] - constructor - · rintro ⟨s, hs, rfl⟩ - show (s.snd : M) * s.fst.prod ∣ y - rw [(unit_associated_one.mul_right s.fst.prod).dvd_iff_dvd_left, one_mul, - ← (normalizedFactors_prod hy).dvd_iff_dvd_right] - exact Multiset.prod_dvd_prod_of_le hs - · rintro (h : x ∣ y) - have hx : x ≠ 0 := by - refine mt (fun hx => ?_) hy - rwa [hx, zero_dvd_iff] at h - obtain ⟨u, hu⟩ := normalizedFactors_prod hx - refine ⟨⟨normalizedFactors x, u⟩, ?_, (mul_comm _ _).trans hu⟩ - exact (dvd_iff_normalizedFactors_le_normalizedFactors hx hy).mp h - -end UniqueFactorizationMonoid - -section Finsupp - -variable [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] -variable [NormalizationMonoid α] [DecidableEq α] - -open UniqueFactorizationMonoid - -/-- This returns the multiset of irreducible factors as a `Finsupp`. -/ -noncomputable def factorization (n : α) : α →₀ ℕ := - Multiset.toFinsupp (normalizedFactors n) - -theorem factorization_eq_count {n p : α} : - factorization n p = Multiset.count p (normalizedFactors n) := by simp [factorization] - -@[simp] -theorem factorization_zero : factorization (0 : α) = 0 := by simp [factorization] - -@[simp] -theorem factorization_one : factorization (1 : α) = 0 := by simp [factorization] - -/-- The support of `factorization n` is exactly the Finset of normalized factors -/ -@[simp] -theorem support_factorization {n : α} : - (factorization n).support = (normalizedFactors n).toFinset := by - simp [factorization, Multiset.toFinsupp_support] - -/-- For nonzero `a` and `b`, the power of `p` in `a * b` is the sum of the powers in `a` and `b` -/ -@[simp] -theorem factorization_mul {a b : α} (ha : a ≠ 0) (hb : b ≠ 0) : - factorization (a * b) = factorization a + factorization b := by - simp [factorization, normalizedFactors_mul ha hb] - -/-- For any `p`, the power of `p` in `x^n` is `n` times the power in `x` -/ -theorem factorization_pow {x : α} {n : ℕ} : factorization (x ^ n) = n • factorization x := by - ext - simp [factorization] - -theorem associated_of_factorization_eq (a b : α) (ha : a ≠ 0) (hb : b ≠ 0) - (h : factorization a = factorization b) : Associated a b := by - simp_rw [factorization, AddEquiv.apply_eq_iff_eq] at h - rwa [associated_iff_normalizedFactors_eq_normalizedFactors ha hb] - -end Finsupp - -open UniqueFactorizationMonoid in -/-- Every non-zero prime ideal in a unique factorization domain contains a prime element. -/ -theorem Ideal.IsPrime.exists_mem_prime_of_ne_bot {R : Type*} [CommSemiring R] [IsDomain R] - [UniqueFactorizationMonoid R] {I : Ideal R} (hI₂ : I.IsPrime) (hI : I ≠ ⊥) : - ∃ x ∈ I, Prime x := by - classical - obtain ⟨a : R, ha₁ : a ∈ I, ha₂ : a ≠ 0⟩ := Submodule.exists_mem_ne_zero_of_ne_bot hI - replace ha₁ : (factors a).prod ∈ I := by - obtain ⟨u : Rˣ, hu : (factors a).prod * u = a⟩ := factors_prod ha₂ - rwa [← hu, mul_unit_mem_iff_mem _ u.isUnit] at ha₁ - obtain ⟨p : R, hp₁ : p ∈ factors a, hp₂ : p ∈ I⟩ := - (hI₂.multiset_prod_mem_iff_exists_mem <| factors a).1 ha₁ - exact ⟨p, hp₂, prime_of_factor p hp₁⟩ - -namespace Nat - -instance instWfDvdMonoid : WfDvdMonoid ℕ where - wf := by - refine RelHomClass.wellFounded - (⟨fun x : ℕ => if x = 0 then (⊤ : ℕ∞) else x, ?_⟩ : DvdNotUnit →r (· < ·)) wellFounded_lt - intro a b h - cases' a with a - · exfalso - revert h - simp [DvdNotUnit] - cases b - · simpa [succ_ne_zero] using WithTop.coe_lt_top (a + 1) - cases' dvd_and_not_dvd_iff.2 h with h1 h2 - simp only [succ_ne_zero, cast_lt, if_false] - refine lt_of_le_of_ne (Nat.le_of_dvd (Nat.succ_pos _) h1) fun con => h2 ?_ - rw [con] - -instance instUniqueFactorizationMonoid : UniqueFactorizationMonoid ℕ where - irreducible_iff_prime := Nat.irreducible_iff_prime - -open UniqueFactorizationMonoid - -lemma factors_eq : ∀ n : ℕ, normalizedFactors n = n.primeFactorsList - | 0 => by simp - | n + 1 => by - rw [← Multiset.rel_eq, ← associated_eq_eq] - apply UniqueFactorizationMonoid.factors_unique irreducible_of_normalized_factor _ - · rw [Multiset.prod_coe, Nat.prod_primeFactorsList n.succ_ne_zero] - exact normalizedFactors_prod n.succ_ne_zero - · intro x hx - rw [Nat.irreducible_iff_prime, ← Nat.prime_iff] - exact Nat.prime_of_mem_primeFactorsList hx - -lemma factors_multiset_prod_of_irreducible {s : Multiset ℕ} (h : ∀ x : ℕ, x ∈ s → Irreducible x) : - normalizedFactors s.prod = s := by - rw [← Multiset.rel_eq, ← associated_eq_eq] - apply UniqueFactorizationMonoid.factors_unique irreducible_of_normalized_factor h - (normalizedFactors_prod _) - rw [Ne, Multiset.prod_eq_zero_iff] - exact fun con ↦ not_irreducible_zero (h 0 con) - -end Nat - -section Ideal - -/-- The ascending chain condition on principal ideals holds in a `WfDvdMonoid` domain. -/ -lemma Ideal.setOf_isPrincipal_wellFoundedOn_gt [CommSemiring α] [WfDvdMonoid α] [IsDomain α] : - {I : Ideal α | I.IsPrincipal}.WellFoundedOn (· > ·) := by - have : {I : Ideal α | I.IsPrincipal} = ((fun a ↦ Ideal.span {a}) '' Set.univ) := by - ext - simp [Submodule.isPrincipal_iff, eq_comm] - rw [this, Set.wellFoundedOn_image, Set.wellFoundedOn_univ] - convert wellFounded_dvdNotUnit (α := α) - ext - exact Ideal.span_singleton_lt_span_singleton - -/-- The ascending chain condition on principal ideals in a domain is sufficient to prove that -the domain is `WfDvdMonoid`. -/ -lemma WfDvdMonoid.of_setOf_isPrincipal_wellFoundedOn_gt [CommSemiring α] [IsDomain α] - (h : {I : Ideal α | I.IsPrincipal}.WellFoundedOn (· > ·)) : - WfDvdMonoid α := by - have : WellFounded (α := {I : Ideal α // I.IsPrincipal}) (· > ·) := h - constructor - convert InvImage.wf (fun a => ⟨Ideal.span ({a} : Set α), _, rfl⟩) this - ext - exact Ideal.span_singleton_lt_span_singleton.symm - -end Ideal - -set_option linter.style.longFile 2100 diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Basic.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Basic.lean new file mode 100644 index 0000000000000..31b9b8426dbed --- /dev/null +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Basic.lean @@ -0,0 +1,473 @@ +/- +Copyright (c) 2018 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Jens Wagemaker, Aaron Anderson +-/ +import Mathlib.Algebra.BigOperators.Associated +import Mathlib.Algebra.Order.Ring.Nat +import Mathlib.Algebra.SMulWithZero +import Mathlib.Data.ENat.Basic +import Mathlib.Data.Multiset.OrderedMonoid +import Mathlib.RingTheory.UniqueFactorizationDomain.Defs + +/-! +# Basic results un unique factorization monoids + +## Main results +* `prime_factors_unique`: the prime factors of an element in a cancellative + commutative monoid with zero (e.g. an integral domain) are unique up to associates +* `UniqueFactorizationMonoid.factors_unique`: the irreducible factors of an element + in a unique factorization monoid (e.g. a UFD) are unique up to associates +* `UniqueFactorizationMonoid.iff_exists_prime_factors`: unique factorization exists iff each nonzero + elements factors into a product of primes +* `UniqueFactorizationMonoid.dvd_of_dvd_mul_left_of_no_prime_factors`: Euclid's lemma: + if `a ∣ b * c` and `a` and `c` have no common prime factors, `a ∣ b`. +* `UniqueFactorizationMonoid.dvd_of_dvd_mul_right_of_no_prime_factors`: Euclid's lemma: + if `a ∣ b * c` and `a` and `b` have no common prime factors, `a ∣ c`. +* `UniqueFactorizationMonoid.exists_reduced_factors`: in a UFM, we can divide out a common factor + to get relatively prime elements. +-/ + + +variable {α : Type*} + +local infixl:50 " ~ᵤ " => Associated + +namespace WfDvdMonoid + +variable [CommMonoidWithZero α] + +open Associates Nat + +theorem of_wfDvdMonoid_associates (_ : WfDvdMonoid (Associates α)) : WfDvdMonoid α := + ⟨(mk_surjective.wellFounded_iff mk_dvdNotUnit_mk_iff.symm).2 wellFounded_dvdNotUnit⟩ + +variable [WfDvdMonoid α] + +instance wfDvdMonoid_associates : WfDvdMonoid (Associates α) := + ⟨(mk_surjective.wellFounded_iff mk_dvdNotUnit_mk_iff.symm).1 wellFounded_dvdNotUnit⟩ + +theorem wellFoundedLT_associates : WellFoundedLT (Associates α) := + ⟨Subrelation.wf dvdNotUnit_of_lt wellFounded_dvdNotUnit⟩ + +@[deprecated wellFoundedLT_associates (since := "2024-09-02")] +theorem wellFounded_associates : WellFounded ((· < ·) : Associates α → Associates α → Prop) := + Subrelation.wf dvdNotUnit_of_lt wellFounded_dvdNotUnit + +end WfDvdMonoid + +theorem WfDvdMonoid.of_wellFoundedLT_associates [CancelCommMonoidWithZero α] + (h : WellFoundedLT (Associates α)) : WfDvdMonoid α := + WfDvdMonoid.of_wfDvdMonoid_associates + ⟨by + convert h.wf + ext + exact Associates.dvdNotUnit_iff_lt⟩ + +@[deprecated WfDvdMonoid.of_wellFoundedLT_associates (since := "2024-09-02")] +theorem WfDvdMonoid.of_wellFounded_associates [CancelCommMonoidWithZero α] + (h : WellFounded ((· < ·) : Associates α → Associates α → Prop)) : WfDvdMonoid α := + WfDvdMonoid.of_wfDvdMonoid_associates + ⟨by + convert h + ext + exact Associates.dvdNotUnit_iff_lt⟩ + +theorem WfDvdMonoid.iff_wellFounded_associates [CancelCommMonoidWithZero α] : + WfDvdMonoid α ↔ WellFoundedLT (Associates α) := + ⟨by apply WfDvdMonoid.wellFoundedLT_associates, WfDvdMonoid.of_wellFoundedLT_associates⟩ + +instance Associates.ufm [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] : + UniqueFactorizationMonoid (Associates α) := + { (WfDvdMonoid.wfDvdMonoid_associates : WfDvdMonoid (Associates α)) with + irreducible_iff_prime := by + rw [← Associates.irreducible_iff_prime_iff] + apply UniqueFactorizationMonoid.irreducible_iff_prime } + +theorem prime_factors_unique [CancelCommMonoidWithZero α] : + ∀ {f g : Multiset α}, + (∀ x ∈ f, Prime x) → (∀ x ∈ g, Prime x) → f.prod ~ᵤ g.prod → Multiset.Rel Associated f g := by + classical + intro f + induction' f using Multiset.induction_on with p f ih + · intros g _ hg h + exact Multiset.rel_zero_left.2 <| + Multiset.eq_zero_of_forall_not_mem fun x hx => + have : IsUnit g.prod := by simpa [associated_one_iff_isUnit] using h.symm + (hg x hx).not_unit <| + isUnit_iff_dvd_one.2 <| (Multiset.dvd_prod hx).trans (isUnit_iff_dvd_one.1 this) + · intros g hf hg hfg + let ⟨b, hbg, hb⟩ := + (exists_associated_mem_of_dvd_prod (hf p (by simp)) fun q hq => hg _ hq) <| + hfg.dvd_iff_dvd_right.1 (show p ∣ (p ::ₘ f).prod by simp) + haveI := Classical.decEq α + rw [← Multiset.cons_erase hbg] + exact + Multiset.Rel.cons hb + (ih (fun q hq => hf _ (by simp [hq])) + (fun {q} (hq : q ∈ g.erase b) => hg q (Multiset.mem_of_mem_erase hq)) + (Associated.of_mul_left + (by rwa [← Multiset.prod_cons, ← Multiset.prod_cons, Multiset.cons_erase hbg]) hb + (hf p (by simp)).ne_zero)) + +namespace UniqueFactorizationMonoid + +variable [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] + +theorem factors_unique {f g : Multiset α} (hf : ∀ x ∈ f, Irreducible x) + (hg : ∀ x ∈ g, Irreducible x) (h : f.prod ~ᵤ g.prod) : Multiset.Rel Associated f g := + prime_factors_unique (fun x hx => UniqueFactorizationMonoid.irreducible_iff_prime.mp (hf x hx)) + (fun x hx => UniqueFactorizationMonoid.irreducible_iff_prime.mp (hg x hx)) h + +end UniqueFactorizationMonoid + +/-- If an irreducible has a prime factorization, + then it is an associate of one of its prime factors. -/ +theorem prime_factors_irreducible [CancelCommMonoidWithZero α] {a : α} {f : Multiset α} + (ha : Irreducible a) (pfa : (∀ b ∈ f, Prime b) ∧ f.prod ~ᵤ a) : ∃ p, a ~ᵤ p ∧ f = {p} := by + haveI := Classical.decEq α + refine @Multiset.induction_on _ + (fun g => (g.prod ~ᵤ a) → (∀ b ∈ g, Prime b) → ∃ p, a ~ᵤ p ∧ g = {p}) f ?_ ?_ pfa.2 pfa.1 + · intro h; exact (ha.not_unit (associated_one_iff_isUnit.1 (Associated.symm h))).elim + · rintro p s _ ⟨u, hu⟩ hs + use p + have hs0 : s = 0 := by + by_contra hs0 + obtain ⟨q, hq⟩ := Multiset.exists_mem_of_ne_zero hs0 + apply (hs q (by simp [hq])).2.1 + refine (ha.isUnit_or_isUnit (?_ : _ = p * ↑u * (s.erase q).prod * _)).resolve_left ?_ + · rw [mul_right_comm _ _ q, mul_assoc, ← Multiset.prod_cons, Multiset.cons_erase hq, ← hu, + mul_comm, mul_comm p _, mul_assoc] + simp + apply mt isUnit_of_mul_isUnit_left (mt isUnit_of_mul_isUnit_left _) + apply (hs p (Multiset.mem_cons_self _ _)).2.1 + simp only [mul_one, Multiset.prod_cons, Multiset.prod_zero, hs0] at * + exact ⟨Associated.symm ⟨u, hu⟩, rfl⟩ + +theorem irreducible_iff_prime_of_exists_unique_irreducible_factors [CancelCommMonoidWithZero α] + (eif : ∀ a : α, a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Irreducible b) ∧ f.prod ~ᵤ a) + (uif : + ∀ f g : Multiset α, + (∀ x ∈ f, Irreducible x) → + (∀ x ∈ g, Irreducible x) → f.prod ~ᵤ g.prod → Multiset.Rel Associated f g) + (p : α) : Irreducible p ↔ Prime p := + letI := Classical.decEq α + ⟨ fun hpi => + ⟨hpi.ne_zero, hpi.1, fun a b ⟨x, hx⟩ => + if hab0 : a * b = 0 then + (eq_zero_or_eq_zero_of_mul_eq_zero hab0).elim (fun ha0 => by simp [ha0]) fun hb0 => by + simp [hb0] + else by + have hx0 : x ≠ 0 := fun hx0 => by simp_all + have ha0 : a ≠ 0 := left_ne_zero_of_mul hab0 + have hb0 : b ≠ 0 := right_ne_zero_of_mul hab0 + cases' eif x hx0 with fx hfx + cases' eif a ha0 with fa hfa + cases' eif b hb0 with fb hfb + have h : Multiset.Rel Associated (p ::ₘ fx) (fa + fb) := by + apply uif + · exact fun i hi => (Multiset.mem_cons.1 hi).elim (fun hip => hip.symm ▸ hpi) (hfx.1 _) + · exact fun i hi => (Multiset.mem_add.1 hi).elim (hfa.1 _) (hfb.1 _) + calc + Multiset.prod (p ::ₘ fx) ~ᵤ a * b := by + rw [hx, Multiset.prod_cons]; exact hfx.2.mul_left _ + _ ~ᵤ fa.prod * fb.prod := hfa.2.symm.mul_mul hfb.2.symm + _ = _ := by rw [Multiset.prod_add] + + exact + let ⟨q, hqf, hq⟩ := Multiset.exists_mem_of_rel_of_mem h (Multiset.mem_cons_self p _) + (Multiset.mem_add.1 hqf).elim + (fun hqa => + Or.inl <| hq.dvd_iff_dvd_left.2 <| hfa.2.dvd_iff_dvd_right.1 (Multiset.dvd_prod hqa)) + fun hqb => + Or.inr <| hq.dvd_iff_dvd_left.2 <| hfb.2.dvd_iff_dvd_right.1 (Multiset.dvd_prod hqb)⟩, + Prime.irreducible⟩ + +namespace UniqueFactorizationMonoid + +variable [CancelCommMonoidWithZero α] +variable [UniqueFactorizationMonoid α] + +@[simp] +theorem factors_one : factors (1 : α) = 0 := by + nontriviality α using factors + rw [← Multiset.rel_zero_right] + refine factors_unique irreducible_of_factor (fun x hx => (Multiset.not_mem_zero x hx).elim) ?_ + rw [Multiset.prod_zero] + exact factors_prod one_ne_zero + +theorem exists_mem_factors_of_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) : + p ∣ a → ∃ q ∈ factors a, p ~ᵤ q := fun ⟨b, hb⟩ => + have hb0 : b ≠ 0 := fun hb0 => by simp_all + have : Multiset.Rel Associated (p ::ₘ factors b) (factors a) := + factors_unique + (fun _ hx => (Multiset.mem_cons.1 hx).elim (fun h => h.symm ▸ hp) (irreducible_of_factor _)) + irreducible_of_factor + (Associated.symm <| + calc + Multiset.prod (factors a) ~ᵤ a := factors_prod ha0 + _ = p * b := hb + _ ~ᵤ Multiset.prod (p ::ₘ factors b) := by + rw [Multiset.prod_cons]; exact (factors_prod hb0).symm.mul_left _ + ) + Multiset.exists_mem_of_rel_of_mem this (by simp) + +theorem exists_mem_factors {x : α} (hx : x ≠ 0) (h : ¬IsUnit x) : ∃ p, p ∈ factors x := by + obtain ⟨p', hp', hp'x⟩ := WfDvdMonoid.exists_irreducible_factor h hx + obtain ⟨p, hp, _⟩ := exists_mem_factors_of_dvd hx hp' hp'x + exact ⟨p, hp⟩ + +open Classical in +theorem factors_mul {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : + Multiset.Rel Associated (factors (x * y)) (factors x + factors y) := by + refine + factors_unique irreducible_of_factor + (fun a ha => + (Multiset.mem_add.mp ha).by_cases (irreducible_of_factor _) (irreducible_of_factor _)) + ((factors_prod (mul_ne_zero hx hy)).trans ?_) + rw [Multiset.prod_add] + exact (Associated.mul_mul (factors_prod hx) (factors_prod hy)).symm + +theorem factors_pow {x : α} (n : ℕ) : + Multiset.Rel Associated (factors (x ^ n)) (n • factors x) := by + match n with + | 0 => rw [zero_smul, pow_zero, factors_one, Multiset.rel_zero_right] + | n+1 => + by_cases h0 : x = 0 + · simp [h0, zero_pow n.succ_ne_zero, smul_zero] + · rw [pow_succ', succ_nsmul'] + refine Multiset.Rel.trans _ (factors_mul h0 (pow_ne_zero n h0)) ?_ + refine Multiset.Rel.add ?_ <| factors_pow n + exact Multiset.rel_refl_of_refl_on fun y _ => Associated.refl _ + +@[simp] +theorem factors_pos (x : α) (hx : x ≠ 0) : 0 < factors x ↔ ¬IsUnit x := by + constructor + · intro h hx + obtain ⟨p, hp⟩ := Multiset.exists_mem_of_ne_zero h.ne' + exact (prime_of_factor _ hp).not_unit (isUnit_of_dvd_unit (dvd_of_mem_factors hp) hx) + · intro h + obtain ⟨p, hp⟩ := exists_mem_factors hx h + exact + bot_lt_iff_ne_bot.mpr + (mt Multiset.eq_zero_iff_forall_not_mem.mp (not_forall.mpr ⟨p, not_not.mpr hp⟩)) + +open Multiset in +theorem factors_pow_count_prod [DecidableEq α] {x : α} (hx : x ≠ 0) : + (∏ p ∈ (factors x).toFinset, p ^ (factors x).count p) ~ᵤ x := + calc + _ = prod (∑ a ∈ toFinset (factors x), count a (factors x) • {a}) := by + simp only [prod_sum, prod_nsmul, prod_singleton] + _ = prod (factors x) := by rw [toFinset_sum_count_nsmul_eq (factors x)] + _ ~ᵤ x := factors_prod hx + +theorem factors_rel_of_associated {a b : α} (h : Associated a b) : + Multiset.Rel Associated (factors a) (factors b) := by + rcases iff_iff_and_or_not_and_not.mp h.eq_zero_iff with (⟨rfl, rfl⟩ | ⟨ha, hb⟩) + · simp + · refine factors_unique irreducible_of_factor irreducible_of_factor ?_ + exact ((factors_prod ha).trans h).trans (factors_prod hb).symm + +end UniqueFactorizationMonoid + +namespace Associates + +attribute [local instance] Associated.setoid + +open Multiset UniqueFactorizationMonoid + +variable [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] + +theorem unique' {p q : Multiset (Associates α)} : + (∀ a ∈ p, Irreducible a) → (∀ a ∈ q, Irreducible a) → p.prod = q.prod → p = q := by + apply Multiset.induction_on_multiset_quot p + apply Multiset.induction_on_multiset_quot q + intro s t hs ht eq + refine Multiset.map_mk_eq_map_mk_of_rel (UniqueFactorizationMonoid.factors_unique ?_ ?_ ?_) + · exact fun a ha => irreducible_mk.1 <| hs _ <| Multiset.mem_map_of_mem _ ha + · exact fun a ha => irreducible_mk.1 <| ht _ <| Multiset.mem_map_of_mem _ ha + have eq' : (Quot.mk Setoid.r : α → Associates α) = Associates.mk := funext quot_mk_eq_mk + rwa [eq', prod_mk, prod_mk, mk_eq_mk_iff_associated] at eq + +theorem prod_le_prod_iff_le [Nontrivial α] {p q : Multiset (Associates α)} + (hp : ∀ a ∈ p, Irreducible a) (hq : ∀ a ∈ q, Irreducible a) : p.prod ≤ q.prod ↔ p ≤ q := by + refine ⟨?_, prod_le_prod⟩ + rintro ⟨c, eqc⟩ + refine Multiset.le_iff_exists_add.2 ⟨factors c, unique' hq (fun x hx ↦ ?_) ?_⟩ + · obtain h | h := Multiset.mem_add.1 hx + · exact hp x h + · exact irreducible_of_factor _ h + · rw [eqc, Multiset.prod_add] + congr + refine associated_iff_eq.mp (factors_prod fun hc => ?_).symm + refine not_irreducible_zero (hq _ ?_) + rw [← prod_eq_zero_iff, eqc, hc, mul_zero] + +end Associates + +section ExistsPrimeFactors + +variable [CancelCommMonoidWithZero α] +variable (pf : ∀ a : α, a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Prime b) ∧ f.prod ~ᵤ a) +include pf + +theorem WfDvdMonoid.of_exists_prime_factors : WfDvdMonoid α := + ⟨by + classical + refine RelHomClass.wellFounded + (RelHom.mk ?_ ?_ : (DvdNotUnit : α → α → Prop) →r ((· < ·) : ℕ∞ → ℕ∞ → Prop)) wellFounded_lt + · intro a + by_cases h : a = 0 + · exact ⊤ + exact ↑(Multiset.card (Classical.choose (pf a h))) + rintro a b ⟨ane0, ⟨c, hc, b_eq⟩⟩ + rw [dif_neg ane0] + by_cases h : b = 0 + · simp [h, lt_top_iff_ne_top] + · rw [dif_neg h, Nat.cast_lt] + have cne0 : c ≠ 0 := by + refine mt (fun con => ?_) h + rw [b_eq, con, mul_zero] + calc + Multiset.card (Classical.choose (pf a ane0)) < + _ + Multiset.card (Classical.choose (pf c cne0)) := + lt_add_of_pos_right _ + (Multiset.card_pos.mpr fun con => hc (associated_one_iff_isUnit.mp ?_)) + _ = Multiset.card (Classical.choose (pf a ane0) + Classical.choose (pf c cne0)) := + (Multiset.card_add _ _).symm + _ = Multiset.card (Classical.choose (pf b h)) := + Multiset.card_eq_card_of_rel + (prime_factors_unique ?_ (Classical.choose_spec (pf _ h)).1 ?_) + + · convert (Classical.choose_spec (pf c cne0)).2.symm + rw [con, Multiset.prod_zero] + · intro x hadd + rw [Multiset.mem_add] at hadd + cases' hadd with h h <;> apply (Classical.choose_spec (pf _ _)).1 _ h <;> assumption + · rw [Multiset.prod_add] + trans a * c + · apply Associated.mul_mul <;> apply (Classical.choose_spec (pf _ _)).2 <;> assumption + · rw [← b_eq] + apply (Classical.choose_spec (pf _ _)).2.symm; assumption⟩ + +theorem irreducible_iff_prime_of_exists_prime_factors {p : α} : Irreducible p ↔ Prime p := by + by_cases hp0 : p = 0 + · simp [hp0] + refine ⟨fun h => ?_, Prime.irreducible⟩ + obtain ⟨f, hf⟩ := pf p hp0 + obtain ⟨q, hq, rfl⟩ := prime_factors_irreducible h hf + rw [hq.prime_iff] + exact hf.1 q (Multiset.mem_singleton_self _) + +theorem UniqueFactorizationMonoid.of_exists_prime_factors : UniqueFactorizationMonoid α := + { WfDvdMonoid.of_exists_prime_factors pf with + irreducible_iff_prime := irreducible_iff_prime_of_exists_prime_factors pf } + +end ExistsPrimeFactors + +theorem UniqueFactorizationMonoid.iff_exists_prime_factors [CancelCommMonoidWithZero α] : + UniqueFactorizationMonoid α ↔ + ∀ a : α, a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Prime b) ∧ f.prod ~ᵤ a := + ⟨fun h => @UniqueFactorizationMonoid.exists_prime_factors _ _ h, + UniqueFactorizationMonoid.of_exists_prime_factors⟩ + +section + +variable {β : Type*} [CancelCommMonoidWithZero α] [CancelCommMonoidWithZero β] + +theorem MulEquiv.uniqueFactorizationMonoid (e : α ≃* β) (hα : UniqueFactorizationMonoid α) : + UniqueFactorizationMonoid β := by + rw [UniqueFactorizationMonoid.iff_exists_prime_factors] at hα ⊢ + intro a ha + obtain ⟨w, hp, u, h⟩ := + hα (e.symm a) fun h => + ha <| by + convert← map_zero e + simp [← h] + exact + ⟨w.map e, fun b hb => + let ⟨c, hc, he⟩ := Multiset.mem_map.1 hb + he ▸ e.prime_iff.1 (hp c hc), + Units.map e.toMonoidHom u, + by + 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 β := + ⟨e.uniqueFactorizationMonoid, e.symm.uniqueFactorizationMonoid⟩ + +end + +namespace UniqueFactorizationMonoid + +theorem of_exists_unique_irreducible_factors [CancelCommMonoidWithZero α] + (eif : ∀ a : α, a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Irreducible b) ∧ f.prod ~ᵤ a) + (uif : + ∀ f g : Multiset α, + (∀ x ∈ f, Irreducible x) → + (∀ x ∈ g, Irreducible x) → f.prod ~ᵤ g.prod → Multiset.Rel Associated f g) : + UniqueFactorizationMonoid α := + UniqueFactorizationMonoid.of_exists_prime_factors + (by + convert eif using 7 + simp_rw [irreducible_iff_prime_of_exists_unique_irreducible_factors eif uif]) + +variable {R : Type*} [CancelCommMonoidWithZero R] [UniqueFactorizationMonoid R] + +theorem isRelPrime_iff_no_prime_factors {a b : R} (ha : a ≠ 0) : + IsRelPrime a b ↔ ∀ ⦃d⦄, d ∣ a → d ∣ b → ¬Prime d := + ⟨fun h _ ha hb ↦ (·.not_unit <| h ha hb), fun h ↦ WfDvdMonoid.isRelPrime_of_no_irreducible_factors + (ha ·.1) fun _ irr ha hb ↦ h ha hb (UniqueFactorizationMonoid.irreducible_iff_prime.mp irr)⟩ + +/-- Euclid's lemma: if `a ∣ b * c` and `a` and `c` have no common prime factors, `a ∣ b`. +Compare `IsCoprime.dvd_of_dvd_mul_left`. -/ +theorem dvd_of_dvd_mul_left_of_no_prime_factors {a b c : R} (ha : a ≠ 0) + (h : ∀ ⦃d⦄, d ∣ a → d ∣ c → ¬Prime d) : a ∣ b * c → a ∣ b := + ((isRelPrime_iff_no_prime_factors ha).mpr h).dvd_of_dvd_mul_right + +/-- Euclid's lemma: if `a ∣ b * c` and `a` and `b` have no common prime factors, `a ∣ c`. +Compare `IsCoprime.dvd_of_dvd_mul_right`. -/ +theorem dvd_of_dvd_mul_right_of_no_prime_factors {a b c : R} (ha : a ≠ 0) + (no_factors : ∀ {d}, d ∣ a → d ∣ b → ¬Prime d) : a ∣ b * c → a ∣ c := by + simpa [mul_comm b c] using dvd_of_dvd_mul_left_of_no_prime_factors ha @no_factors + +/-- If `a ≠ 0, b` are elements of a unique factorization domain, then dividing +out their common factor `c'` gives `a'` and `b'` with no factors in common. -/ +theorem exists_reduced_factors : + ∀ a ≠ (0 : R), ∀ b, + ∃ a' b' c', IsRelPrime a' b' ∧ c' * a' = a ∧ c' * b' = b := by + intro a + refine induction_on_prime a ?_ ?_ ?_ + · intros + contradiction + · intro a a_unit _ b + use a, b, 1 + constructor + · intro p p_dvd_a _ + exact isUnit_of_dvd_unit p_dvd_a a_unit + · simp + · intro a p a_ne_zero p_prime ih_a pa_ne_zero b + by_cases h : p ∣ b + · rcases h with ⟨b, rfl⟩ + obtain ⟨a', b', c', no_factor, ha', hb'⟩ := ih_a a_ne_zero b + refine ⟨a', b', p * c', @no_factor, ?_, ?_⟩ + · rw [mul_assoc, ha'] + · rw [mul_assoc, hb'] + · obtain ⟨a', b', c', coprime, rfl, rfl⟩ := ih_a a_ne_zero b + refine ⟨p * a', b', c', ?_, mul_left_comm _ _ _, rfl⟩ + intro q q_dvd_pa' q_dvd_b' + cases' p_prime.left_dvd_or_dvd_right_of_dvd_mul q_dvd_pa' with p_dvd_q q_dvd_a' + · have : p ∣ c' * b' := dvd_mul_of_dvd_right (p_dvd_q.trans q_dvd_b') _ + contradiction + exact coprime q_dvd_a' q_dvd_b' + +theorem exists_reduced_factors' (a b : R) (hb : b ≠ 0) : + ∃ a' b' c', IsRelPrime a' b' ∧ c' * a' = a ∧ c' * b' = b := + let ⟨b', a', c', no_factor, hb, ha⟩ := exists_reduced_factors b hb a + ⟨a', b', c', fun _ hpb hpa => no_factor hpa hpb, ha, hb⟩ + +@[deprecated (since := "2024-09-21")] alias pow_right_injective := pow_injective_of_not_isUnit +@[deprecated (since := "2024-09-21")] alias pow_eq_pow_iff := pow_inj_of_not_isUnit + +end UniqueFactorizationMonoid diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Defs.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Defs.lean new file mode 100644 index 0000000000000..121cf52e3a622 --- /dev/null +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Defs.lean @@ -0,0 +1,205 @@ +/- +Copyright (c) 2018 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Jens Wagemaker, Aaron Anderson +-/ +import Mathlib.Algebra.Associated.Basic +import Mathlib.Algebra.BigOperators.Group.Multiset +import Mathlib.Algebra.Group.Submonoid.Membership +import Mathlib.Order.WellFounded + +/-! +# Unique factorization + +## Main Definitions +* `WfDvdMonoid` holds for `Monoid`s for which a strict divisibility relation is + well-founded. +* `UniqueFactorizationMonoid` holds for `WfDvdMonoid`s where + `Irreducible` is equivalent to `Prime` +-/ + +assert_not_exists Field +assert_not_exists Finsupp +assert_not_exists Ideal + +variable {α : Type*} + +local infixl:50 " ~ᵤ " => Associated + +/-- Well-foundedness of the strict version of ∣, which is equivalent to the descending chain +condition on divisibility and to the ascending chain condition on +principal ideals in an integral domain. + -/ +abbrev WfDvdMonoid (α : Type*) [CommMonoidWithZero α] : Prop := + IsWellFounded α DvdNotUnit + +theorem wellFounded_dvdNotUnit {α : Type*} [CommMonoidWithZero α] [h : WfDvdMonoid α] : + WellFounded (DvdNotUnit (α := α)) := + h.wf + +namespace WfDvdMonoid + +variable [CommMonoidWithZero α] + +open Associates Nat + +variable [WfDvdMonoid α] + +theorem exists_irreducible_factor {a : α} (ha : ¬IsUnit a) (ha0 : a ≠ 0) : + ∃ i, Irreducible i ∧ i ∣ a := + let ⟨b, hs, hr⟩ := wellFounded_dvdNotUnit.has_min { b | b ∣ a ∧ ¬IsUnit b } ⟨a, dvd_rfl, ha⟩ + ⟨b, + ⟨hs.2, fun c d he => + let h := dvd_trans ⟨d, he⟩ hs.1 + or_iff_not_imp_left.2 fun hc => + of_not_not fun hd => hr c ⟨h, hc⟩ ⟨ne_zero_of_dvd_ne_zero ha0 h, d, hd, he⟩⟩, + hs.1⟩ + +@[elab_as_elim] +theorem induction_on_irreducible {P : α → Prop} (a : α) (h0 : P 0) (hu : ∀ u : α, IsUnit u → P u) + (hi : ∀ a i : α, a ≠ 0 → Irreducible i → P a → P (i * a)) : P a := + haveI := Classical.dec + wellFounded_dvdNotUnit.fix + (fun a ih => + if ha0 : a = 0 then ha0.substr h0 + else + if hau : IsUnit a then hu a hau + else + let ⟨i, hii, b, hb⟩ := exists_irreducible_factor hau ha0 + let hb0 : b ≠ 0 := ne_zero_of_dvd_ne_zero ha0 ⟨i, mul_comm i b ▸ hb⟩ + hb.symm ▸ hi b i hb0 hii <| ih b ⟨hb0, i, hii.1, mul_comm i b ▸ hb⟩) + a + +theorem exists_factors (a : α) : + a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Irreducible b) ∧ Associated f.prod a := + induction_on_irreducible a (fun h => (h rfl).elim) + (fun _ hu _ => ⟨0, fun _ h => False.elim (Multiset.not_mem_zero _ h), hu.unit, one_mul _⟩) + fun a i ha0 hi ih _ => + let ⟨s, hs⟩ := ih ha0 + ⟨i ::ₘ s, fun b H => (Multiset.mem_cons.1 H).elim (fun h => h.symm ▸ hi) (hs.1 b), by + rw [s.prod_cons i] + exact hs.2.mul_left i⟩ + +theorem not_unit_iff_exists_factors_eq (a : α) (hn0 : a ≠ 0) : + ¬IsUnit a ↔ ∃ f : Multiset α, (∀ b ∈ f, Irreducible b) ∧ f.prod = a ∧ f ≠ ∅ := + ⟨fun hnu => by + obtain ⟨f, hi, u, rfl⟩ := exists_factors a hn0 + obtain ⟨b, h⟩ := Multiset.exists_mem_of_ne_zero fun h : f = 0 => hnu <| by simp [h] + classical + refine ⟨(f.erase b).cons (b * u), fun a ha => ?_, ?_, Multiset.cons_ne_zero⟩ + · obtain rfl | ha := Multiset.mem_cons.1 ha + exacts [Associated.irreducible ⟨u, rfl⟩ (hi b h), hi a (Multiset.mem_of_mem_erase ha)] + · rw [Multiset.prod_cons, mul_comm b, mul_assoc, Multiset.prod_erase h, mul_comm], + fun ⟨_, hi, he, hne⟩ => + let ⟨b, h⟩ := Multiset.exists_mem_of_ne_zero hne + not_isUnit_of_not_isUnit_dvd (hi b h).not_unit <| he ▸ Multiset.dvd_prod h⟩ + +theorem isRelPrime_of_no_irreducible_factors {x y : α} (nonzero : ¬(x = 0 ∧ y = 0)) + (H : ∀ z : α, Irreducible z → z ∣ x → ¬z ∣ y) : IsRelPrime x y := + isRelPrime_of_no_nonunits_factors nonzero fun _z znu znz zx zy ↦ + have ⟨i, h1, h2⟩ := exists_irreducible_factor znu znz + H i h1 (h2.trans zx) (h2.trans zy) + +end WfDvdMonoid + +section Prio + +-- set_option default_priority 100 + +-- see Note [default priority] +/-- unique factorization monoids. + +These are defined as `CancelCommMonoidWithZero`s with well-founded strict divisibility +relations, but this is equivalent to more familiar definitions: + +Each element (except zero) is uniquely represented as a multiset of irreducible factors. +Uniqueness is only up to associated elements. + +Each element (except zero) is non-uniquely represented as a multiset +of prime factors. + +To define a UFD using the definition in terms of multisets +of irreducible factors, use the definition `of_exists_unique_irreducible_factors` + +To define a UFD using the definition in terms of multisets +of prime factors, use the definition `of_exists_prime_factors` + +-/ +class UniqueFactorizationMonoid (α : Type*) [CancelCommMonoidWithZero α] extends + IsWellFounded α DvdNotUnit : Prop where + protected irreducible_iff_prime : ∀ {a : α}, Irreducible a ↔ Prime a + +instance (priority := 100) ufm_of_decomposition_of_wfDvdMonoid + [CancelCommMonoidWithZero α] [WfDvdMonoid α] [DecompositionMonoid α] : + UniqueFactorizationMonoid α := + { ‹WfDvdMonoid α› with irreducible_iff_prime := irreducible_iff_prime } + +@[deprecated ufm_of_decomposition_of_wfDvdMonoid (since := "2024-02-12")] +theorem ufm_of_gcd_of_wfDvdMonoid [CancelCommMonoidWithZero α] [WfDvdMonoid α] + [DecompositionMonoid α] : UniqueFactorizationMonoid α := + ufm_of_decomposition_of_wfDvdMonoid + +end Prio + +namespace UniqueFactorizationMonoid + +variable [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] + +theorem exists_prime_factors (a : α) : + a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Prime b) ∧ f.prod ~ᵤ a := by + simp_rw [← UniqueFactorizationMonoid.irreducible_iff_prime] + apply WfDvdMonoid.exists_factors a + +lemma exists_prime_iff : + (∃ (p : α), Prime p) ↔ ∃ (x : α), x ≠ 0 ∧ ¬ IsUnit x := by + refine ⟨fun ⟨p, hp⟩ ↦ ⟨p, hp.ne_zero, hp.not_unit⟩, fun ⟨x, hx₀, hxu⟩ ↦ ?_⟩ + obtain ⟨f, hf, -⟩ := WfDvdMonoid.exists_irreducible_factor hxu hx₀ + exact ⟨f, UniqueFactorizationMonoid.irreducible_iff_prime.mp hf⟩ + +@[elab_as_elim] +theorem induction_on_prime {P : α → Prop} (a : α) (h₁ : P 0) (h₂ : ∀ x : α, IsUnit x → P x) + (h₃ : ∀ a p : α, a ≠ 0 → Prime p → P a → P (p * a)) : P a := by + simp_rw [← UniqueFactorizationMonoid.irreducible_iff_prime] at h₃ + exact WfDvdMonoid.induction_on_irreducible a h₁ h₂ h₃ + +instance : DecompositionMonoid α where + primal a := by + obtain rfl | ha := eq_or_ne a 0; · exact isPrimal_zero + obtain ⟨f, hf, u, rfl⟩ := exists_prime_factors a ha + exact ((Submonoid.isPrimal α).multiset_prod_mem f (hf · ·|>.isPrimal)).mul u.isUnit.isPrimal + +end UniqueFactorizationMonoid + +namespace UniqueFactorizationMonoid + +variable [CancelCommMonoidWithZero α] +variable [UniqueFactorizationMonoid α] + +open Classical in +/-- Noncomputably determines the multiset of prime factors. -/ +noncomputable def factors (a : α) : Multiset α := + if h : a = 0 then 0 else Classical.choose (UniqueFactorizationMonoid.exists_prime_factors a h) + +theorem factors_prod {a : α} (ane0 : a ≠ 0) : Associated (factors a).prod a := by + rw [factors, dif_neg ane0] + exact (Classical.choose_spec (exists_prime_factors a ane0)).2 + +@[simp] +theorem factors_zero : factors (0 : α) = 0 := by simp [factors] + +theorem ne_zero_of_mem_factors {p a : α} (h : p ∈ factors a) : a ≠ 0 := by + rintro rfl + simp at h + +theorem dvd_of_mem_factors {p a : α} (h : p ∈ factors a) : p ∣ a := + dvd_trans (Multiset.dvd_prod h) (Associated.dvd (factors_prod (ne_zero_of_mem_factors h))) + +theorem prime_of_factor {a : α} (x : α) (hx : x ∈ factors a) : Prime x := by + have ane0 := ne_zero_of_mem_factors hx + rw [factors, dif_neg ane0] at hx + exact (Classical.choose_spec (UniqueFactorizationMonoid.exists_prime_factors a ane0)).1 x hx + +theorem irreducible_of_factor {a : α} : ∀ x : α, x ∈ factors a → Irreducible x := fun x h => + (prime_of_factor x h).irreducible + +end UniqueFactorizationMonoid diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/FactorSet.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/FactorSet.lean new file mode 100644 index 0000000000000..2dd1d63b54515 --- /dev/null +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/FactorSet.lean @@ -0,0 +1,647 @@ +/- +Copyright (c) 2018 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Jens Wagemaker, Aaron Anderson +-/ +import Mathlib.Data.Finsupp.Multiset +import Mathlib.RingTheory.UniqueFactorizationDomain.Basic +import Mathlib.Tactic.Ring + +/-! +# Set of factors + +## Main definitions +* `Associates.FactorSet`: multiset of factors of an element, unique up to propositional equality. +* `Associates.factors`: determine the `FactorSet` for a given element. + +## TODO +* set up the complete lattice structure on `FactorSet`. + +-/ + +variable {α : Type*} + +local infixl:50 " ~ᵤ " => Associated + +namespace Associates + +open UniqueFactorizationMonoid Associated Multiset + +variable [CancelCommMonoidWithZero α] + +/-- `FactorSet α` representation elements of unique factorization domain as multisets. +`Multiset α` produced by `normalizedFactors` are only unique up to associated elements, while the +multisets in `FactorSet α` are unique by equality and restricted to irreducible elements. This +gives us a representation of each element as a unique multisets (or the added ⊤ for 0), which has a +complete lattice structure. Infimum is the greatest common divisor and supremum is the least common +multiple. +-/ +abbrev FactorSet.{u} (α : Type u) [CancelCommMonoidWithZero α] : Type u := + WithTop (Multiset { a : Associates α // Irreducible a }) + +attribute [local instance] Associated.setoid + +theorem FactorSet.coe_add {a b : Multiset { a : Associates α // Irreducible a }} : + (↑(a + b) : FactorSet α) = a + b := by norm_cast + +theorem FactorSet.sup_add_inf_eq_add [DecidableEq (Associates α)] : + ∀ a b : FactorSet α, a ⊔ b + a ⊓ b = a + b + | ⊤, b => show ⊤ ⊔ b + ⊤ ⊓ b = ⊤ + b by simp + | a, ⊤ => show a ⊔ ⊤ + a ⊓ ⊤ = a + ⊤ by simp + | WithTop.some a, WithTop.some b => + show (a : FactorSet α) ⊔ b + (a : FactorSet α) ⊓ b = a + b by + rw [← WithTop.coe_sup, ← WithTop.coe_inf, ← WithTop.coe_add, ← WithTop.coe_add, + WithTop.coe_eq_coe] + exact Multiset.union_add_inter _ _ + +/-- Evaluates the product of a `FactorSet` to be the product of the corresponding multiset, + or `0` if there is none. -/ +def FactorSet.prod : FactorSet α → Associates α + | ⊤ => 0 + | WithTop.some s => (s.map (↑)).prod + +@[simp] +theorem prod_top : (⊤ : FactorSet α).prod = 0 := + rfl + +@[simp] +theorem prod_coe {s : Multiset { a : Associates α // Irreducible a }} : + FactorSet.prod (s : FactorSet α) = (s.map (↑)).prod := + rfl + +@[simp] +theorem prod_add : ∀ a b : FactorSet α, (a + b).prod = a.prod * b.prod + | ⊤, b => show (⊤ + b).prod = (⊤ : FactorSet α).prod * b.prod by simp + | a, ⊤ => show (a + ⊤).prod = a.prod * (⊤ : FactorSet α).prod by simp + | WithTop.some a, WithTop.some b => by + rw [← FactorSet.coe_add, prod_coe, prod_coe, prod_coe, Multiset.map_add, Multiset.prod_add] + +@[gcongr] +theorem prod_mono : ∀ {a b : FactorSet α}, a ≤ b → a.prod ≤ b.prod + | ⊤, b, h => by + have : b = ⊤ := top_unique h + rw [this, prod_top] + | a, ⊤, _ => show a.prod ≤ (⊤ : FactorSet α).prod by simp + | WithTop.some _, WithTop.some _, h => + prod_le_prod <| Multiset.map_le_map <| WithTop.coe_le_coe.1 <| h + +theorem FactorSet.prod_eq_zero_iff [Nontrivial α] (p : FactorSet α) : p.prod = 0 ↔ p = ⊤ := by + unfold FactorSet at p + induction p -- TODO: `induction_eliminator` doesn't work with `abbrev` + · simp only [eq_self_iff_true, Associates.prod_top] + · rw [prod_coe, Multiset.prod_eq_zero_iff, Multiset.mem_map, eq_false WithTop.coe_ne_top, + iff_false, not_exists] + exact fun a => not_and_of_not_right _ a.prop.ne_zero + +section count + +variable [DecidableEq (Associates α)] + +/-- `bcount p s` is the multiplicity of `p` in the FactorSet `s` (with bundled `p`)-/ +def bcount (p : { a : Associates α // Irreducible a }) : + FactorSet α → ℕ + | ⊤ => 0 + | WithTop.some s => s.count p + +variable [∀ p : Associates α, Decidable (Irreducible p)] {p : Associates α} + +/-- `count p s` is the multiplicity of the irreducible `p` in the FactorSet `s`. + +If `p` is not irreducible, `count p s` is defined to be `0`. -/ +def count (p : Associates α) : FactorSet α → ℕ := + if hp : Irreducible p then bcount ⟨p, hp⟩ else 0 + +@[simp] +theorem count_some (hp : Irreducible p) (s : Multiset _) : + count p (WithTop.some s) = s.count ⟨p, hp⟩ := by + simp only [count, dif_pos hp, bcount] + +@[simp] +theorem count_zero (hp : Irreducible p) : count p (0 : FactorSet α) = 0 := by + simp only [count, dif_pos hp, bcount, Multiset.count_zero] + +theorem count_reducible (hp : ¬Irreducible p) : count p = 0 := dif_neg hp + +end count + +section Mem + +/-- membership in a FactorSet (bundled version) -/ +def BfactorSetMem : { a : Associates α // Irreducible a } → FactorSet α → Prop + | _, ⊤ => True + | p, some l => p ∈ l + +/-- `FactorSetMem p s` is the predicate that the irreducible `p` is a member of +`s : FactorSet α`. + +If `p` is not irreducible, `p` is not a member of any `FactorSet`. -/ +def FactorSetMem (s : FactorSet α) (p : Associates α) : Prop := + letI : Decidable (Irreducible p) := Classical.dec _ + if hp : Irreducible p then BfactorSetMem ⟨p, hp⟩ s else False + +instance : Membership (Associates α) (FactorSet α) := + ⟨FactorSetMem⟩ + +@[simp] +theorem factorSetMem_eq_mem (p : Associates α) (s : FactorSet α) : FactorSetMem s p = (p ∈ s) := + rfl + +theorem mem_factorSet_top {p : Associates α} {hp : Irreducible p} : p ∈ (⊤ : FactorSet α) := by + dsimp only [Membership.mem]; dsimp only [FactorSetMem]; split_ifs; exact trivial + +theorem mem_factorSet_some {p : Associates α} {hp : Irreducible p} + {l : Multiset { a : Associates α // Irreducible a }} : + p ∈ (l : FactorSet α) ↔ Subtype.mk p hp ∈ l := by + dsimp only [Membership.mem]; dsimp only [FactorSetMem]; split_ifs; rfl + +theorem reducible_not_mem_factorSet {p : Associates α} (hp : ¬Irreducible p) (s : FactorSet α) : + ¬p ∈ s := fun h ↦ by + rwa [← factorSetMem_eq_mem, FactorSetMem, dif_neg hp] at h + +theorem irreducible_of_mem_factorSet {p : Associates α} {s : FactorSet α} (h : p ∈ s) : + Irreducible p := + by_contra fun hp ↦ reducible_not_mem_factorSet hp s h + +end Mem + +variable [UniqueFactorizationMonoid α] + +theorem FactorSet.unique [Nontrivial α] {p q : FactorSet α} (h : p.prod = q.prod) : p = q := by + -- TODO: `induction_eliminator` doesn't work with `abbrev` + unfold FactorSet at p q + induction p <;> induction q + · rfl + · rw [eq_comm, ← FactorSet.prod_eq_zero_iff, ← h, Associates.prod_top] + · rw [← FactorSet.prod_eq_zero_iff, h, Associates.prod_top] + · congr 1 + rw [← Multiset.map_eq_map Subtype.coe_injective] + apply unique' _ _ h <;> + · intro a ha + obtain ⟨⟨a', irred⟩, -, rfl⟩ := Multiset.mem_map.mp ha + rwa [Subtype.coe_mk] + +/-- This returns the multiset of irreducible factors as a `FactorSet`, + a multiset of irreducible associates `WithTop`. -/ +noncomputable def factors' (a : α) : Multiset { a : Associates α // Irreducible a } := + (factors a).pmap (fun a ha => ⟨Associates.mk a, irreducible_mk.2 ha⟩) irreducible_of_factor + +@[simp] +theorem map_subtype_coe_factors' {a : α} : + (factors' a).map (↑) = (factors a).map Associates.mk := by + simp [factors', Multiset.map_pmap, Multiset.pmap_eq_map] + +theorem factors'_cong {a b : α} (h : a ~ᵤ b) : factors' a = factors' b := by + obtain rfl | hb := eq_or_ne b 0 + · rw [associated_zero_iff_eq_zero] at h + rw [h] + have ha : a ≠ 0 := by + contrapose! hb with ha + rw [← associated_zero_iff_eq_zero, ← ha] + exact h.symm + rw [← Multiset.map_eq_map Subtype.coe_injective, map_subtype_coe_factors', + map_subtype_coe_factors', ← rel_associated_iff_map_eq_map] + exact + factors_unique irreducible_of_factor irreducible_of_factor + ((factors_prod ha).trans <| h.trans <| (factors_prod hb).symm) + +/-- This returns the multiset of irreducible factors of an associate as a `FactorSet`, + a multiset of irreducible associates `WithTop`. -/ +noncomputable def factors (a : Associates α) : FactorSet α := by + classical refine if h : a = 0 then ⊤ else Quotient.hrecOn a (fun x _ => factors' x) ?_ h + intro a b hab + apply Function.hfunext + · have : a ~ᵤ 0 ↔ b ~ᵤ 0 := Iff.intro (fun ha0 => hab.symm.trans ha0) fun hb0 => hab.trans hb0 + simp only [associated_zero_iff_eq_zero] at this + simp only [quotient_mk_eq_mk, this, mk_eq_zero] + exact fun ha hb _ => heq_of_eq <| congr_arg some <| factors'_cong hab + +@[simp] +theorem factors_zero : (0 : Associates α).factors = ⊤ := + dif_pos rfl + +@[deprecated (since := "2024-03-16")] alias factors_0 := factors_zero + +@[simp] +theorem factors_mk (a : α) (h : a ≠ 0) : (Associates.mk a).factors = factors' a := by + classical + apply dif_neg + apply mt mk_eq_zero.1 h + +@[simp] +theorem factors_prod (a : Associates α) : a.factors.prod = a := by + rcases Associates.mk_surjective a with ⟨a, rfl⟩ + rcases eq_or_ne a 0 with rfl | ha + · simp + · simp [ha, prod_mk, mk_eq_mk_iff_associated, UniqueFactorizationMonoid.factors_prod, + -Quotient.eq] + +@[simp] +theorem prod_factors [Nontrivial α] (s : FactorSet α) : s.prod.factors = s := + FactorSet.unique <| factors_prod _ + +@[nontriviality] +theorem factors_subsingleton [Subsingleton α] {a : Associates α} : a.factors = ⊤ := by + have : Subsingleton (Associates α) := inferInstance + convert factors_zero + +theorem factors_eq_top_iff_zero {a : Associates α} : a.factors = ⊤ ↔ a = 0 := by + nontriviality α + exact ⟨fun h ↦ by rwa [← factors_prod a, FactorSet.prod_eq_zero_iff], fun h ↦ h ▸ factors_zero⟩ + +@[deprecated (since := "2024-04-16")] alias factors_eq_none_iff_zero := factors_eq_top_iff_zero + +theorem factors_eq_some_iff_ne_zero {a : Associates α} : + (∃ s : Multiset { p : Associates α // Irreducible p }, a.factors = s) ↔ a ≠ 0 := by + simp_rw [@eq_comm _ a.factors, ← WithTop.ne_top_iff_exists] + exact factors_eq_top_iff_zero.not + +theorem eq_of_factors_eq_factors {a b : Associates α} (h : a.factors = b.factors) : a = b := by + have : a.factors.prod = b.factors.prod := by rw [h] + rwa [factors_prod, factors_prod] at this + +theorem eq_of_prod_eq_prod [Nontrivial α] {a b : FactorSet α} (h : a.prod = b.prod) : a = b := by + have : a.prod.factors = b.prod.factors := by rw [h] + rwa [prod_factors, prod_factors] at this + +@[simp] +theorem factors_mul (a b : Associates α) : (a * b).factors = a.factors + b.factors := by + nontriviality α + refine eq_of_prod_eq_prod <| eq_of_factors_eq_factors ?_ + rw [prod_add, factors_prod, factors_prod, factors_prod] + +@[gcongr] +theorem factors_mono : ∀ {a b : Associates α}, a ≤ b → a.factors ≤ b.factors + | s, t, ⟨d, eq⟩ => by rw [eq, factors_mul]; exact le_add_of_nonneg_right bot_le + +@[simp] +theorem factors_le {a b : Associates α} : a.factors ≤ b.factors ↔ a ≤ b := by + refine ⟨fun h ↦ ?_, factors_mono⟩ + have : a.factors.prod ≤ b.factors.prod := prod_mono h + rwa [factors_prod, factors_prod] at this + +section count + +variable [DecidableEq (Associates α)] [∀ p : Associates α, Decidable (Irreducible p)] + +theorem eq_factors_of_eq_counts {a b : Associates α} (ha : a ≠ 0) (hb : b ≠ 0) + (h : ∀ p : Associates α, Irreducible p → p.count a.factors = p.count b.factors) : + a.factors = b.factors := by + obtain ⟨sa, h_sa⟩ := factors_eq_some_iff_ne_zero.mpr ha + obtain ⟨sb, h_sb⟩ := factors_eq_some_iff_ne_zero.mpr hb + rw [h_sa, h_sb] at h ⊢ + rw [WithTop.coe_eq_coe] + have h_count : ∀ (p : Associates α) (hp : Irreducible p), + sa.count ⟨p, hp⟩ = sb.count ⟨p, hp⟩ := by + intro p hp + rw [← count_some, ← count_some, h p hp] + apply Multiset.toFinsupp.injective + ext ⟨p, hp⟩ + rw [Multiset.toFinsupp_apply, Multiset.toFinsupp_apply, h_count p hp] + +theorem eq_of_eq_counts {a b : Associates α} (ha : a ≠ 0) (hb : b ≠ 0) + (h : ∀ p : Associates α, Irreducible p → p.count a.factors = p.count b.factors) : a = b := + eq_of_factors_eq_factors (eq_factors_of_eq_counts ha hb h) + +theorem count_le_count_of_factors_le {a b p : Associates α} (hb : b ≠ 0) (hp : Irreducible p) + (h : a.factors ≤ b.factors) : p.count a.factors ≤ p.count b.factors := by + by_cases ha : a = 0 + · simp_all + obtain ⟨sa, h_sa⟩ := factors_eq_some_iff_ne_zero.mpr ha + obtain ⟨sb, h_sb⟩ := factors_eq_some_iff_ne_zero.mpr hb + rw [h_sa, h_sb] at h ⊢ + rw [count_some hp, count_some hp]; rw [WithTop.coe_le_coe] at h + exact Multiset.count_le_of_le _ h + +theorem count_le_count_of_le {a b p : Associates α} (hb : b ≠ 0) (hp : Irreducible p) (h : a ≤ b) : + p.count a.factors ≤ p.count b.factors := + count_le_count_of_factors_le hb hp <| factors_mono h + +end count + +theorem prod_le [Nontrivial α] {a b : FactorSet α} : a.prod ≤ b.prod ↔ a ≤ b := by + refine ⟨fun h ↦ ?_, prod_mono⟩ + have : a.prod.factors ≤ b.prod.factors := factors_mono h + rwa [prod_factors, prod_factors] at this + +open Classical in +noncomputable instance : Max (Associates α) := + ⟨fun a b => (a.factors ⊔ b.factors).prod⟩ + +open Classical in +noncomputable instance : Min (Associates α) := + ⟨fun a b => (a.factors ⊓ b.factors).prod⟩ + +open Classical in +noncomputable instance : Lattice (Associates α) := + { Associates.instPartialOrder with + sup := (· ⊔ ·) + inf := (· ⊓ ·) + sup_le := fun _ _ c hac hbc => + factors_prod c ▸ prod_mono (sup_le (factors_mono hac) (factors_mono hbc)) + le_sup_left := fun a _ => le_trans (le_of_eq (factors_prod a).symm) <| prod_mono <| le_sup_left + le_sup_right := fun _ b => + le_trans (le_of_eq (factors_prod b).symm) <| prod_mono <| le_sup_right + le_inf := fun a _ _ hac hbc => + factors_prod a ▸ prod_mono (le_inf (factors_mono hac) (factors_mono hbc)) + inf_le_left := fun a _ => le_trans (prod_mono inf_le_left) (le_of_eq (factors_prod a)) + inf_le_right := fun _ b => le_trans (prod_mono inf_le_right) (le_of_eq (factors_prod b)) } + +open Classical in +theorem sup_mul_inf (a b : Associates α) : (a ⊔ b) * (a ⊓ b) = a * b := + show (a.factors ⊔ b.factors).prod * (a.factors ⊓ b.factors).prod = a * b by + nontriviality α + refine eq_of_factors_eq_factors ?_ + rw [← prod_add, prod_factors, factors_mul, FactorSet.sup_add_inf_eq_add] + +theorem dvd_of_mem_factors {a p : Associates α} (hm : p ∈ factors a) : + p ∣ a := by + rcases eq_or_ne a 0 with rfl | ha0 + · exact dvd_zero p + obtain ⟨a0, nza, ha'⟩ := exists_non_zero_rep ha0 + rw [← Associates.factors_prod a] + rw [← ha', factors_mk a0 nza] at hm ⊢ + rw [prod_coe] + apply Multiset.dvd_prod; apply Multiset.mem_map.mpr + exact ⟨⟨p, irreducible_of_mem_factorSet hm⟩, mem_factorSet_some.mp hm, rfl⟩ + +theorem dvd_of_mem_factors' {a : α} {p : Associates α} {hp : Irreducible p} {hz : a ≠ 0} + (h_mem : Subtype.mk p hp ∈ factors' a) : p ∣ Associates.mk a := by + haveI := Classical.decEq (Associates α) + apply dvd_of_mem_factors + rw [factors_mk _ hz] + apply mem_factorSet_some.2 h_mem + +theorem mem_factors'_of_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) (hd : p ∣ a) : + Subtype.mk (Associates.mk p) (irreducible_mk.2 hp) ∈ factors' a := by + obtain ⟨q, hq, hpq⟩ := exists_mem_factors_of_dvd ha0 hp hd + apply Multiset.mem_pmap.mpr; use q; use hq + exact Subtype.eq (Eq.symm (mk_eq_mk_iff_associated.mpr hpq)) + +theorem mem_factors'_iff_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) : + Subtype.mk (Associates.mk p) (irreducible_mk.2 hp) ∈ factors' a ↔ p ∣ a := by + constructor + · rw [← mk_dvd_mk] + apply dvd_of_mem_factors' + apply ha0 + · apply mem_factors'_of_dvd ha0 hp + +theorem mem_factors_of_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) (hd : p ∣ a) : + Associates.mk p ∈ factors (Associates.mk a) := by + rw [factors_mk _ ha0] + exact mem_factorSet_some.mpr (mem_factors'_of_dvd ha0 hp hd) + +theorem mem_factors_iff_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) : + Associates.mk p ∈ factors (Associates.mk a) ↔ p ∣ a := by + constructor + · rw [← mk_dvd_mk] + apply dvd_of_mem_factors + · apply mem_factors_of_dvd ha0 hp + +open Classical in +theorem exists_prime_dvd_of_not_inf_one {a b : α} (ha : a ≠ 0) (hb : b ≠ 0) + (h : Associates.mk a ⊓ Associates.mk b ≠ 1) : ∃ p : α, Prime p ∧ p ∣ a ∧ p ∣ b := by + have hz : factors (Associates.mk a) ⊓ factors (Associates.mk b) ≠ 0 := by + contrapose! h with hf + change (factors (Associates.mk a) ⊓ factors (Associates.mk b)).prod = 1 + rw [hf] + exact Multiset.prod_zero + rw [factors_mk a ha, factors_mk b hb, ← WithTop.coe_inf] at hz + obtain ⟨⟨p0, p0_irr⟩, p0_mem⟩ := Multiset.exists_mem_of_ne_zero ((mt WithTop.coe_eq_coe.mpr) hz) + rw [Multiset.inf_eq_inter] at p0_mem + obtain ⟨p, rfl⟩ : ∃ p, Associates.mk p = p0 := Quot.exists_rep p0 + refine ⟨p, ?_, ?_, ?_⟩ + · rw [← UniqueFactorizationMonoid.irreducible_iff_prime, ← irreducible_mk] + exact p0_irr + · apply dvd_of_mk_le_mk + apply dvd_of_mem_factors' (Multiset.mem_inter.mp p0_mem).left + apply ha + · apply dvd_of_mk_le_mk + apply dvd_of_mem_factors' (Multiset.mem_inter.mp p0_mem).right + apply hb + +theorem coprime_iff_inf_one {a b : α} (ha0 : a ≠ 0) (hb0 : b ≠ 0) : + Associates.mk a ⊓ Associates.mk b = 1 ↔ ∀ {d : α}, d ∣ a → d ∣ b → ¬Prime d := by + constructor + · intro hg p ha hb hp + refine (Associates.prime_mk.mpr hp).not_unit (isUnit_of_dvd_one ?_) + rw [← hg] + exact le_inf (mk_le_mk_of_dvd ha) (mk_le_mk_of_dvd hb) + · contrapose + intro hg hc + obtain ⟨p, hp, hpa, hpb⟩ := exists_prime_dvd_of_not_inf_one ha0 hb0 hg + exact hc hpa hpb hp + +theorem factors_self [Nontrivial α] {p : Associates α} (hp : Irreducible p) : + p.factors = WithTop.some {⟨p, hp⟩} := + eq_of_prod_eq_prod + (by rw [factors_prod, FactorSet.prod.eq_def]; dsimp; rw [prod_singleton]) + +theorem factors_prime_pow [Nontrivial α] {p : Associates α} (hp : Irreducible p) (k : ℕ) : + factors (p ^ k) = WithTop.some (Multiset.replicate k ⟨p, hp⟩) := + eq_of_prod_eq_prod + (by + rw [Associates.factors_prod, FactorSet.prod.eq_def] + dsimp; rw [Multiset.map_replicate, Multiset.prod_replicate, Subtype.coe_mk]) + +theorem prime_pow_le_iff_le_bcount [DecidableEq (Associates α)] {m p : Associates α} + (h₁ : m ≠ 0) (h₂ : Irreducible p) {k : ℕ} : p ^ k ≤ m ↔ k ≤ bcount ⟨p, h₂⟩ m.factors := by + rcases Associates.exists_non_zero_rep h₁ with ⟨m, hm, rfl⟩ + have := nontrivial_of_ne _ _ hm + rw [bcount.eq_def, factors_mk, Multiset.le_count_iff_replicate_le, ← factors_le, + factors_prime_pow, factors_mk, WithTop.coe_le_coe] <;> assumption + +@[simp] +theorem factors_one [Nontrivial α] : factors (1 : Associates α) = 0 := by + apply eq_of_prod_eq_prod + rw [Associates.factors_prod] + exact Multiset.prod_zero + +@[simp] +theorem pow_factors [Nontrivial α] {a : Associates α} {k : ℕ} : + (a ^ k).factors = k • a.factors := by + induction' k with n h + · rw [zero_nsmul, pow_zero] + exact factors_one + · rw [pow_succ, succ_nsmul, factors_mul, h] + +section count + +variable [DecidableEq (Associates α)] [∀ p : Associates α, Decidable (Irreducible p)] + +theorem prime_pow_dvd_iff_le {m p : Associates α} (h₁ : m ≠ 0) (h₂ : Irreducible p) {k : ℕ} : + p ^ k ≤ m ↔ k ≤ count p m.factors := by + rw [count, dif_pos h₂, prime_pow_le_iff_le_bcount h₁] + +theorem le_of_count_ne_zero {m p : Associates α} (h0 : m ≠ 0) (hp : Irreducible p) : + count p m.factors ≠ 0 → p ≤ m := by + nontriviality α + rw [← pos_iff_ne_zero] + intro h + rw [← pow_one p] + apply (prime_pow_dvd_iff_le h0 hp).2 + simpa only + +theorem count_ne_zero_iff_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) : + (Associates.mk p).count (Associates.mk a).factors ≠ 0 ↔ p ∣ a := by + nontriviality α + rw [← Associates.mk_le_mk_iff_dvd] + refine + ⟨fun h => + Associates.le_of_count_ne_zero (Associates.mk_ne_zero.mpr ha0) + (Associates.irreducible_mk.mpr hp) h, + fun h => ?_⟩ + rw [← pow_one (Associates.mk p), + Associates.prime_pow_dvd_iff_le (Associates.mk_ne_zero.mpr ha0) + (Associates.irreducible_mk.mpr hp)] at h + exact (zero_lt_one.trans_le h).ne' + +theorem count_self [Nontrivial α] {p : Associates α} + (hp : Irreducible p) : p.count p.factors = 1 := by + simp [factors_self hp, Associates.count_some hp] + +theorem count_eq_zero_of_ne {p q : Associates α} (hp : Irreducible p) + (hq : Irreducible q) (h : p ≠ q) : p.count q.factors = 0 := + not_ne_iff.mp fun h' ↦ h <| associated_iff_eq.mp <| hp.associated_of_dvd hq <| + le_of_count_ne_zero hq.ne_zero hp h' + +theorem count_mul {a : Associates α} (ha : a ≠ 0) {b : Associates α} + (hb : b ≠ 0) {p : Associates α} (hp : Irreducible p) : + count p (factors (a * b)) = count p a.factors + count p b.factors := by + obtain ⟨a0, nza, rfl⟩ := exists_non_zero_rep ha + obtain ⟨b0, nzb, rfl⟩ := exists_non_zero_rep hb + rw [factors_mul, factors_mk a0 nza, factors_mk b0 nzb, ← FactorSet.coe_add, count_some hp, + Multiset.count_add, count_some hp, count_some hp] + +theorem count_of_coprime {a : Associates α} (ha : a ≠ 0) + {b : Associates α} (hb : b ≠ 0) (hab : ∀ d, d ∣ a → d ∣ b → ¬Prime d) {p : Associates α} + (hp : Irreducible p) : count p a.factors = 0 ∨ count p b.factors = 0 := by + rw [or_iff_not_imp_left, ← Ne] + intro hca + contrapose! hab with hcb + exact ⟨p, le_of_count_ne_zero ha hp hca, le_of_count_ne_zero hb hp hcb, + UniqueFactorizationMonoid.irreducible_iff_prime.mp hp⟩ + +theorem count_mul_of_coprime {a : Associates α} {b : Associates α} + (hb : b ≠ 0) {p : Associates α} (hp : Irreducible p) (hab : ∀ d, d ∣ a → d ∣ b → ¬Prime d) : + count p a.factors = 0 ∨ count p a.factors = count p (a * b).factors := by + by_cases ha : a = 0 + · simp [ha] + cases' count_of_coprime ha hb hab hp with hz hb0; · tauto + apply Or.intro_right + rw [count_mul ha hb hp, hb0, add_zero] + +theorem count_mul_of_coprime' {a b : Associates α} {p : Associates α} + (hp : Irreducible p) (hab : ∀ d, d ∣ a → d ∣ b → ¬Prime d) : + count p (a * b).factors = count p a.factors ∨ count p (a * b).factors = count p b.factors := by + by_cases ha : a = 0 + · simp [ha] + by_cases hb : b = 0 + · simp [hb] + rw [count_mul ha hb hp] + cases' count_of_coprime ha hb hab hp with ha0 hb0 + · apply Or.intro_right + rw [ha0, zero_add] + · apply Or.intro_left + rw [hb0, add_zero] + +theorem dvd_count_of_dvd_count_mul {a b : Associates α} (hb : b ≠ 0) + {p : Associates α} (hp : Irreducible p) (hab : ∀ d, d ∣ a → d ∣ b → ¬Prime d) {k : ℕ} + (habk : k ∣ count p (a * b).factors) : k ∣ count p a.factors := by + by_cases ha : a = 0 + · simpa [*] using habk + cases' count_of_coprime ha hb hab hp with hz h + · rw [hz] + exact dvd_zero k + · rw [count_mul ha hb hp, h] at habk + exact habk + +theorem count_pow [Nontrivial α] {a : Associates α} (ha : a ≠ 0) + {p : Associates α} (hp : Irreducible p) (k : ℕ) : + count p (a ^ k).factors = k * count p a.factors := by + induction' k with n h + · rw [pow_zero, factors_one, zero_mul, count_zero hp] + · rw [pow_succ', count_mul ha (pow_ne_zero _ ha) hp, h] + ring + +theorem dvd_count_pow [Nontrivial α] {a : Associates α} (ha : a ≠ 0) + {p : Associates α} (hp : Irreducible p) (k : ℕ) : k ∣ count p (a ^ k).factors := by + rw [count_pow ha hp] + apply dvd_mul_right + +theorem is_pow_of_dvd_count {a : Associates α} + (ha : a ≠ 0) {k : ℕ} (hk : ∀ p : Associates α, Irreducible p → k ∣ count p a.factors) : + ∃ b : Associates α, a = b ^ k := by + nontriviality α + obtain ⟨a0, hz, rfl⟩ := exists_non_zero_rep ha + rw [factors_mk a0 hz] at hk + have hk' : ∀ p, p ∈ factors' a0 → k ∣ (factors' a0).count p := by + rintro p - + have pp : p = ⟨p.val, p.2⟩ := by simp only [Subtype.coe_eta] + rw [pp, ← count_some p.2] + exact hk p.val p.2 + obtain ⟨u, hu⟩ := Multiset.exists_smul_of_dvd_count _ hk' + use FactorSet.prod (u : FactorSet α) + apply eq_of_factors_eq_factors + rw [pow_factors, prod_factors, factors_mk a0 hz, hu] + exact WithBot.coe_nsmul u k + +/-- The only divisors of prime powers are prime powers. See `eq_pow_find_of_dvd_irreducible_pow` +for an explicit expression as a p-power (without using `count`). -/ +theorem eq_pow_count_factors_of_dvd_pow {p a : Associates α} + (hp : Irreducible p) {n : ℕ} (h : a ∣ p ^ n) : a = p ^ p.count a.factors := by + nontriviality α + have hph := pow_ne_zero n hp.ne_zero + have ha := ne_zero_of_dvd_ne_zero hph h + apply eq_of_eq_counts ha (pow_ne_zero _ hp.ne_zero) + have eq_zero_of_ne : ∀ q : Associates α, Irreducible q → q ≠ p → _ = 0 := fun q hq h' => + Nat.eq_zero_of_le_zero <| by + convert count_le_count_of_le hph hq h + symm + rw [count_pow hp.ne_zero hq, count_eq_zero_of_ne hq hp h', mul_zero] + intro q hq + rw [count_pow hp.ne_zero hq] + by_cases h : q = p + · rw [h, count_self hp, mul_one] + · rw [count_eq_zero_of_ne hq hp h, mul_zero, eq_zero_of_ne q hq h] + +theorem count_factors_eq_find_of_dvd_pow {a p : Associates α} + (hp : Irreducible p) [∀ n : ℕ, Decidable (a ∣ p ^ n)] {n : ℕ} (h : a ∣ p ^ n) : + @Nat.find (fun n => a ∣ p ^ n) _ ⟨n, h⟩ = p.count a.factors := by + apply le_antisymm + · refine Nat.find_le ⟨1, ?_⟩ + rw [mul_one] + symm + exact eq_pow_count_factors_of_dvd_pow hp h + · have hph := pow_ne_zero (@Nat.find (fun n => a ∣ p ^ n) _ ⟨n, h⟩) hp.ne_zero + cases' subsingleton_or_nontrivial α with hα hα + · simp [eq_iff_true_of_subsingleton] at hph + convert count_le_count_of_le hph hp (@Nat.find_spec (fun n => a ∣ p ^ n) _ ⟨n, h⟩) + rw [count_pow hp.ne_zero hp, count_self hp, mul_one] + +end count + +theorem eq_pow_of_mul_eq_pow {a b c : Associates α} (ha : a ≠ 0) (hb : b ≠ 0) + (hab : ∀ d, d ∣ a → d ∣ b → ¬Prime d) {k : ℕ} (h : a * b = c ^ k) : + ∃ d : Associates α, a = d ^ k := by + classical + nontriviality α + by_cases hk0 : k = 0 + · use 1 + rw [hk0, pow_zero] at h ⊢ + apply (mul_eq_one.1 h).1 + · refine is_pow_of_dvd_count ha fun p hp ↦ ?_ + apply dvd_count_of_dvd_count_mul hb hp hab + rw [h] + apply dvd_count_pow _ hp + rintro rfl + rw [zero_pow hk0] at h + cases mul_eq_zero.mp h <;> contradiction + +/-- The only divisors of prime powers are prime powers. -/ +theorem eq_pow_find_of_dvd_irreducible_pow {a p : Associates α} (hp : Irreducible p) + [∀ n : ℕ, Decidable (a ∣ p ^ n)] {n : ℕ} (h : a ∣ p ^ n) : + a = p ^ @Nat.find (fun n => a ∣ p ^ n) _ ⟨n, h⟩ := by + classical rw [count_factors_eq_find_of_dvd_pow hp, ← eq_pow_count_factors_of_dvd_pow hp h] + exact h + +end Associates diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Finite.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Finite.lean new file mode 100644 index 0000000000000..bd5cd5c1e9819 --- /dev/null +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Finite.lean @@ -0,0 +1,53 @@ +/- +Copyright (c) 2018 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Jens Wagemaker, Aaron Anderson +-/ +import Mathlib.RingTheory.UniqueFactorizationDomain.NormalizedFactors + +/-! +# Finiteness of divisors + +## Main results +* `UniqueFactorizationMonoid.fintypeSubtypeDvd`: elements of a UFM with finitely many units have + finitely many divisors. +-/ + +variable {α : Type*} + +local infixl:50 " ~ᵤ " => Associated + +namespace UniqueFactorizationMonoid + +/-- If `y` is a nonzero element of a unique factorization monoid with finitely +many units (e.g. `ℤ`, `Ideal (ring_of_integers K)`), it has finitely many divisors. -/ +noncomputable def fintypeSubtypeDvd {M : Type*} [CancelCommMonoidWithZero M] + [UniqueFactorizationMonoid M] [Fintype Mˣ] (y : M) (hy : y ≠ 0) : Fintype { x // x ∣ y } := by + haveI : Nontrivial M := ⟨⟨y, 0, hy⟩⟩ + haveI : NormalizationMonoid M := UniqueFactorizationMonoid.normalizationMonoid + haveI := Classical.decEq M + haveI := Classical.decEq (Associates M) + -- We'll show `fun (u : Mˣ) (f ⊆ factors y) ↦ u * Π f` is injective + -- and has image exactly the divisors of `y`. + refine + Fintype.ofFinset + (((normalizedFactors y).powerset.toFinset ×ˢ (Finset.univ : Finset Mˣ)).image fun s => + (s.snd : M) * s.fst.prod) + fun x => ?_ + simp only [exists_prop, Finset.mem_image, Finset.mem_product, Finset.mem_univ, and_true, + Multiset.mem_toFinset, Multiset.mem_powerset, exists_eq_right, Multiset.mem_map] + constructor + · rintro ⟨s, hs, rfl⟩ + show (s.snd : M) * s.fst.prod ∣ y + rw [(unit_associated_one.mul_right s.fst.prod).dvd_iff_dvd_left, one_mul, + ← (normalizedFactors_prod hy).dvd_iff_dvd_right] + exact Multiset.prod_dvd_prod_of_le hs + · rintro (h : x ∣ y) + have hx : x ≠ 0 := by + refine mt (fun hx => ?_) hy + rwa [hx, zero_dvd_iff] at h + obtain ⟨u, hu⟩ := normalizedFactors_prod hx + refine ⟨⟨normalizedFactors x, u⟩, ?_, (mul_comm _ _).trans hu⟩ + exact (dvd_iff_normalizedFactors_le_normalizedFactors hx hy).mp h + +end UniqueFactorizationMonoid diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Finsupp.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Finsupp.lean new file mode 100644 index 0000000000000..a901b3c83b1a3 --- /dev/null +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Finsupp.lean @@ -0,0 +1,62 @@ +/- +Copyright (c) 2018 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Jens Wagemaker, Aaron Anderson +-/ +import Mathlib.Data.Finsupp.Multiset +import Mathlib.RingTheory.UniqueFactorizationDomain.NormalizedFactors + +/-! +# Factors as finsupp + +## Main definitions +* `UniqueFactorizationMonoid.factorization`: the multiset of irreducible factors as a `Finsupp`. +-/ + +variable {α : Type*} + +local infixl:50 " ~ᵤ " => Associated + +section Finsupp + +variable [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] +variable [NormalizationMonoid α] [DecidableEq α] + +open UniqueFactorizationMonoid + +/-- This returns the multiset of irreducible factors as a `Finsupp`. -/ +noncomputable def factorization (n : α) : α →₀ ℕ := + Multiset.toFinsupp (normalizedFactors n) + +theorem factorization_eq_count {n p : α} : + factorization n p = Multiset.count p (normalizedFactors n) := by simp [factorization] + +@[simp] +theorem factorization_zero : factorization (0 : α) = 0 := by simp [factorization] + +@[simp] +theorem factorization_one : factorization (1 : α) = 0 := by simp [factorization] + +/-- The support of `factorization n` is exactly the Finset of normalized factors -/ +@[simp] +theorem support_factorization {n : α} : + (factorization n).support = (normalizedFactors n).toFinset := by + simp [factorization, Multiset.toFinsupp_support] + +/-- For nonzero `a` and `b`, the power of `p` in `a * b` is the sum of the powers in `a` and `b` -/ +@[simp] +theorem factorization_mul {a b : α} (ha : a ≠ 0) (hb : b ≠ 0) : + factorization (a * b) = factorization a + factorization b := by + simp [factorization, normalizedFactors_mul ha hb] + +/-- For any `p`, the power of `p` in `x^n` is `n` times the power in `x` -/ +theorem factorization_pow {x : α} {n : ℕ} : factorization (x ^ n) = n • factorization x := by + ext + simp [factorization] + +theorem associated_of_factorization_eq (a b : α) (ha : a ≠ 0) (hb : b ≠ 0) + (h : factorization a = factorization b) : Associated a b := by + simp_rw [factorization, AddEquiv.apply_eq_iff_eq] at h + rwa [associated_iff_normalizedFactors_eq_normalizedFactors ha hb] + +end Finsupp diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/GCDMonoid.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/GCDMonoid.lean new file mode 100644 index 0000000000000..74940f8f2ed5f --- /dev/null +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/GCDMonoid.lean @@ -0,0 +1,73 @@ +/- +Copyright (c) 2018 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Jens Wagemaker, Aaron Anderson +-/ +import Mathlib.RingTheory.UniqueFactorizationDomain.FactorSet +import Mathlib.RingTheory.UniqueFactorizationDomain.NormalizedFactors + +/-! +# Building GCD out of unique factorization + +## Main results +* `UniqueFactorizationMonoid.toGCDMonoid`: choose a GCD monoid structure given unique factorization. +-/ + +variable {α : Type*} + +local infixl:50 " ~ᵤ " => Associated + +section + +open Associates UniqueFactorizationMonoid + +/-- `toGCDMonoid` constructs a GCD monoid out of a unique factorization domain. -/ +noncomputable def UniqueFactorizationMonoid.toGCDMonoid (α : Type*) [CancelCommMonoidWithZero α] + [UniqueFactorizationMonoid α] : GCDMonoid α where + gcd a b := Quot.out (Associates.mk a ⊓ Associates.mk b : Associates α) + lcm a b := Quot.out (Associates.mk a ⊔ Associates.mk b : Associates α) + gcd_dvd_left a b := by + rw [← mk_dvd_mk, Associates.quot_out, congr_fun₂ dvd_eq_le] + exact inf_le_left + gcd_dvd_right a b := by + rw [← mk_dvd_mk, Associates.quot_out, congr_fun₂ dvd_eq_le] + exact inf_le_right + dvd_gcd {a b c} hac hab := by + rw [← mk_dvd_mk, Associates.quot_out, congr_fun₂ dvd_eq_le, le_inf_iff, + mk_le_mk_iff_dvd, mk_le_mk_iff_dvd] + exact ⟨hac, hab⟩ + lcm_zero_left a := by simp + lcm_zero_right a := by simp + gcd_mul_lcm a b := by + rw [← mk_eq_mk_iff_associated, ← Associates.mk_mul_mk, ← associated_iff_eq, Associates.quot_out, + Associates.quot_out, mul_comm, sup_mul_inf, Associates.mk_mul_mk] + +/-- `toNormalizedGCDMonoid` constructs a GCD monoid out of a normalization on a + unique factorization domain. -/ +noncomputable def UniqueFactorizationMonoid.toNormalizedGCDMonoid (α : Type*) + [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] [NormalizationMonoid α] : + NormalizedGCDMonoid α := + { ‹NormalizationMonoid α› with + gcd := fun a b => (Associates.mk a ⊓ Associates.mk b).out + lcm := fun a b => (Associates.mk a ⊔ Associates.mk b).out + gcd_dvd_left := fun a b => (out_dvd_iff a (Associates.mk a ⊓ Associates.mk b)).2 <| inf_le_left + gcd_dvd_right := fun a b => + (out_dvd_iff b (Associates.mk a ⊓ Associates.mk b)).2 <| inf_le_right + dvd_gcd := fun {a} {b} {c} hac hab => + show a ∣ (Associates.mk c ⊓ Associates.mk b).out by + rw [dvd_out_iff, le_inf_iff, mk_le_mk_iff_dvd, mk_le_mk_iff_dvd] + exact ⟨hac, hab⟩ + lcm_zero_left := fun a => show (⊤ ⊔ Associates.mk a).out = 0 by simp + lcm_zero_right := fun a => show (Associates.mk a ⊔ ⊤).out = 0 by simp + gcd_mul_lcm := fun a b => by + rw [← out_mul, mul_comm, sup_mul_inf, mk_mul_mk, out_mk] + exact normalize_associated (a * b) + normalize_gcd := fun a b => by apply normalize_out _ + normalize_lcm := fun a b => by apply normalize_out _ } + +instance (α) [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] : + Nonempty (NormalizedGCDMonoid α) := by + letI := UniqueFactorizationMonoid.normalizationMonoid (α := α) + classical exact ⟨UniqueFactorizationMonoid.toNormalizedGCDMonoid α⟩ + +end diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Ideal.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Ideal.lean new file mode 100644 index 0000000000000..76a9fb5c09941 --- /dev/null +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Ideal.lean @@ -0,0 +1,58 @@ +/- +Copyright (c) 2018 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Jens Wagemaker, Aaron Anderson +-/ +import Mathlib.RingTheory.Ideal.Operations +import Mathlib.RingTheory.UniqueFactorizationDomain.Defs + +/-! +# Unique factorization and ascending chain condition on ideals + +## Main results +* `Ideal.setOf_isPrincipal_wellFoundedOn_gt`, `WfDvdMonoid.of_setOf_isPrincipal_wellFoundedOn_gt` + in a domain, well-foundedness of the strict version of ∣ is equivalent to the ascending + chain condition on principal ideals. +-/ + +variable {α : Type*} + +open UniqueFactorizationMonoid in +/-- Every non-zero prime ideal in a unique factorization domain contains a prime element. -/ +theorem Ideal.IsPrime.exists_mem_prime_of_ne_bot {R : Type*} [CommSemiring R] [IsDomain R] + [UniqueFactorizationMonoid R] {I : Ideal R} (hI₂ : I.IsPrime) (hI : I ≠ ⊥) : + ∃ x ∈ I, Prime x := by + classical + obtain ⟨a : R, ha₁ : a ∈ I, ha₂ : a ≠ 0⟩ := Submodule.exists_mem_ne_zero_of_ne_bot hI + replace ha₁ : (factors a).prod ∈ I := by + obtain ⟨u : Rˣ, hu : (factors a).prod * u = a⟩ := factors_prod ha₂ + rwa [← hu, mul_unit_mem_iff_mem _ u.isUnit] at ha₁ + obtain ⟨p : R, hp₁ : p ∈ factors a, hp₂ : p ∈ I⟩ := + (hI₂.multiset_prod_mem_iff_exists_mem <| factors a).1 ha₁ + exact ⟨p, hp₂, prime_of_factor p hp₁⟩ + +section Ideal + +/-- The ascending chain condition on principal ideals holds in a `WfDvdMonoid` domain. -/ +lemma Ideal.setOf_isPrincipal_wellFoundedOn_gt [CommSemiring α] [WfDvdMonoid α] [IsDomain α] : + {I : Ideal α | I.IsPrincipal}.WellFoundedOn (· > ·) := by + have : {I : Ideal α | I.IsPrincipal} = ((fun a ↦ Ideal.span {a}) '' Set.univ) := by + ext + simp [Submodule.isPrincipal_iff, eq_comm] + rw [this, Set.wellFoundedOn_image, Set.wellFoundedOn_univ] + convert wellFounded_dvdNotUnit (α := α) + ext + exact Ideal.span_singleton_lt_span_singleton + +/-- The ascending chain condition on principal ideals in a domain is sufficient to prove that +the domain is `WfDvdMonoid`. -/ +lemma WfDvdMonoid.of_setOf_isPrincipal_wellFoundedOn_gt [CommSemiring α] [IsDomain α] + (h : {I : Ideal α | I.IsPrincipal}.WellFoundedOn (· > ·)) : + WfDvdMonoid α := by + have : WellFounded (α := {I : Ideal α // I.IsPrincipal}) (· > ·) := h + constructor + convert InvImage.wf (fun a => ⟨Ideal.span ({a} : Set α), _, rfl⟩) this + ext + exact Ideal.span_singleton_lt_span_singleton.symm + +end Ideal diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicative.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicative.lean new file mode 100644 index 0000000000000..a8a04d3a9ad0f --- /dev/null +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicative.lean @@ -0,0 +1,163 @@ +/- +Copyright (c) 2018 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Jens Wagemaker, Aaron Anderson +-/ +import Mathlib.RingTheory.UniqueFactorizationDomain.NormalizedFactors + +/-! +# Multiplicative maps on unique factorization domains + +## Main results +* `UniqueFactorizationMonoid.induction_on_coprime`: if `P` holds for `0`, units and powers of + primes, and `P x ∧ P y` for coprime `x, y` implies `P (x * y)`, then `P` holds on all `a : α`. +* `UniqueFactorizationMonoid.multiplicative_of_coprime`: if `f` maps `p ^ i` to `(f p) ^ i` for + primes `p`, and `f` is multiplicative on coprime elements, then `f` is multiplicative everywhere. +-/ + + +variable {α : Type*} + +namespace UniqueFactorizationMonoid + +variable {R : Type*} [CancelCommMonoidWithZero R] [UniqueFactorizationMonoid R] + +section Multiplicative + +variable [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] +variable {β : Type*} [CancelCommMonoidWithZero β] + +theorem prime_pow_coprime_prod_of_coprime_insert [DecidableEq α] {s : Finset α} (i : α → ℕ) (p : α) + (hps : p ∉ s) (is_prime : ∀ q ∈ insert p s, Prime q) + (is_coprime : ∀ᵉ (q ∈ insert p s) (q' ∈ insert p s), q ∣ q' → q = q') : + IsRelPrime (p ^ i p) (∏ p' ∈ s, p' ^ i p') := by + have hp := is_prime _ (Finset.mem_insert_self _ _) + refine (isRelPrime_iff_no_prime_factors <| pow_ne_zero _ hp.ne_zero).mpr ?_ + intro d hdp hdprod hd + apply hps + replace hdp := hd.dvd_of_dvd_pow hdp + obtain ⟨q, q_mem', hdq⟩ := hd.exists_mem_multiset_dvd hdprod + obtain ⟨q, q_mem, rfl⟩ := Multiset.mem_map.mp q_mem' + replace hdq := hd.dvd_of_dvd_pow hdq + have : p ∣ q := dvd_trans (hd.irreducible.dvd_symm hp.irreducible hdp) hdq + convert q_mem using 0 + rw [Finset.mem_val, + is_coprime _ (Finset.mem_insert_self p s) _ (Finset.mem_insert_of_mem q_mem) this] + +/-- If `P` holds for units and powers of primes, +and `P x ∧ P y` for coprime `x, y` implies `P (x * y)`, +then `P` holds on a product of powers of distinct primes. -/ +@[elab_as_elim] +theorem induction_on_prime_power {P : α → Prop} (s : Finset α) (i : α → ℕ) + (is_prime : ∀ p ∈ s, Prime p) (is_coprime : ∀ᵉ (p ∈ s) (q ∈ s), p ∣ q → p = q) + (h1 : ∀ {x}, IsUnit x → P x) (hpr : ∀ {p} (i : ℕ), Prime p → P (p ^ i)) + (hcp : ∀ {x y}, IsRelPrime x y → P x → P y → P (x * y)) : + P (∏ p ∈ s, p ^ i p) := by + letI := Classical.decEq α + induction' s using Finset.induction_on with p f' hpf' ih + · simpa using h1 isUnit_one + rw [Finset.prod_insert hpf'] + exact + hcp (prime_pow_coprime_prod_of_coprime_insert i p hpf' is_prime is_coprime) + (hpr (i p) (is_prime _ (Finset.mem_insert_self _ _))) + (ih (fun q hq => is_prime _ (Finset.mem_insert_of_mem hq)) fun q hq q' hq' => + is_coprime _ (Finset.mem_insert_of_mem hq) _ (Finset.mem_insert_of_mem hq')) + +/-- If `P` holds for `0`, units and powers of primes, +and `P x ∧ P y` for coprime `x, y` implies `P (x * y)`, +then `P` holds on all `a : α`. -/ +@[elab_as_elim] +theorem induction_on_coprime {P : α → Prop} (a : α) (h0 : P 0) (h1 : ∀ {x}, IsUnit x → P x) + (hpr : ∀ {p} (i : ℕ), Prime p → P (p ^ i)) + (hcp : ∀ {x y}, IsRelPrime x y → P x → P y → P (x * y)) : P a := by + letI := Classical.decEq α + have P_of_associated : ∀ {x y}, Associated x y → P x → P y := by + rintro x y ⟨u, rfl⟩ hx + exact hcp (fun p _ hpx => isUnit_of_dvd_unit hpx u.isUnit) hx (h1 u.isUnit) + by_cases ha0 : a = 0 + · rwa [ha0] + haveI : Nontrivial α := ⟨⟨_, _, ha0⟩⟩ + letI : NormalizationMonoid α := UniqueFactorizationMonoid.normalizationMonoid + refine P_of_associated (normalizedFactors_prod ha0) ?_ + rw [← (normalizedFactors a).map_id, Finset.prod_multiset_map_count] + refine induction_on_prime_power _ _ ?_ ?_ @h1 @hpr @hcp <;> simp only [Multiset.mem_toFinset] + · apply prime_of_normalized_factor + · apply normalizedFactors_eq_of_dvd + +/-- If `f` maps `p ^ i` to `(f p) ^ i` for primes `p`, and `f` +is multiplicative on coprime elements, then `f` is multiplicative on all products of primes. -/ +theorem multiplicative_prime_power {f : α → β} (s : Finset α) (i j : α → ℕ) + (is_prime : ∀ p ∈ s, Prime p) (is_coprime : ∀ᵉ (p ∈ s) (q ∈ s), p ∣ q → p = q) + (h1 : ∀ {x y}, IsUnit y → f (x * y) = f x * f y) + (hpr : ∀ {p} (i : ℕ), Prime p → f (p ^ i) = f p ^ i) + (hcp : ∀ {x y}, IsRelPrime x y → f (x * y) = f x * f y) : + f (∏ p ∈ s, p ^ (i p + j p)) = f (∏ p ∈ s, p ^ i p) * f (∏ p ∈ s, p ^ j p) := by + letI := Classical.decEq α + induction' s using Finset.induction_on with p s hps ih + · simpa using h1 isUnit_one + have hpr_p := is_prime _ (Finset.mem_insert_self _ _) + have hpr_s : ∀ p ∈ s, Prime p := fun p hp => is_prime _ (Finset.mem_insert_of_mem hp) + have hcp_p := fun i => prime_pow_coprime_prod_of_coprime_insert i p hps is_prime is_coprime + have hcp_s : ∀ᵉ (p ∈ s) (q ∈ s), p ∣ q → p = q := fun p hp q hq => + is_coprime p (Finset.mem_insert_of_mem hp) q (Finset.mem_insert_of_mem hq) + rw [Finset.prod_insert hps, Finset.prod_insert hps, Finset.prod_insert hps, hcp (hcp_p _), + hpr _ hpr_p, hcp (hcp_p _), hpr _ hpr_p, hcp (hcp_p (fun p => i p + j p)), hpr _ hpr_p, + ih hpr_s hcp_s, pow_add, mul_assoc, mul_left_comm (f p ^ j p), mul_assoc] + +/-- If `f` maps `p ^ i` to `(f p) ^ i` for primes `p`, and `f` +is multiplicative on coprime elements, then `f` is multiplicative everywhere. -/ +theorem multiplicative_of_coprime (f : α → β) (a b : α) (h0 : f 0 = 0) + (h1 : ∀ {x y}, IsUnit y → f (x * y) = f x * f y) + (hpr : ∀ {p} (i : ℕ), Prime p → f (p ^ i) = f p ^ i) + (hcp : ∀ {x y}, IsRelPrime x y → f (x * y) = f x * f y) : + f (a * b) = f a * f b := by + letI := Classical.decEq α + by_cases ha0 : a = 0 + · rw [ha0, zero_mul, h0, zero_mul] + by_cases hb0 : b = 0 + · rw [hb0, mul_zero, h0, mul_zero] + by_cases hf1 : f 1 = 0 + · calc + f (a * b) = f (a * b * 1) := by rw [mul_one] + _ = 0 := by simp only [h1 isUnit_one, hf1, mul_zero] + _ = f a * f (b * 1) := by simp only [h1 isUnit_one, hf1, mul_zero] + _ = f a * f b := by rw [mul_one] + haveI : Nontrivial α := ⟨⟨_, _, ha0⟩⟩ + letI : NormalizationMonoid α := UniqueFactorizationMonoid.normalizationMonoid + suffices + f (∏ p ∈ (normalizedFactors a).toFinset ∪ (normalizedFactors b).toFinset, + p ^ ((normalizedFactors a).count p + (normalizedFactors b).count p)) = + f (∏ p ∈ (normalizedFactors a).toFinset ∪ (normalizedFactors b).toFinset, + p ^ (normalizedFactors a).count p) * + f (∏ p ∈ (normalizedFactors a).toFinset ∪ (normalizedFactors b).toFinset, + p ^ (normalizedFactors b).count p) by + obtain ⟨ua, a_eq⟩ := normalizedFactors_prod ha0 + obtain ⟨ub, b_eq⟩ := normalizedFactors_prod hb0 + rw [← a_eq, ← b_eq, mul_right_comm (Multiset.prod (normalizedFactors a)) ua + (Multiset.prod (normalizedFactors b) * ub), h1 ua.isUnit, h1 ub.isUnit, h1 ua.isUnit, ← + mul_assoc, h1 ub.isUnit, mul_right_comm _ (f ua), ← mul_assoc] + congr + rw [← (normalizedFactors a).map_id, ← (normalizedFactors b).map_id, + Finset.prod_multiset_map_count, Finset.prod_multiset_map_count, + Finset.prod_subset (Finset.subset_union_left (s₂ := (normalizedFactors b).toFinset)), + Finset.prod_subset (Finset.subset_union_right (s₂ := (normalizedFactors b).toFinset)), ← + Finset.prod_mul_distrib] + · simp_rw [id, ← pow_add, this] + all_goals simp only [Multiset.mem_toFinset] + · intro p _ hpb + simp [hpb] + · intro p _ hpa + simp [hpa] + refine multiplicative_prime_power _ _ _ ?_ ?_ @h1 @hpr @hcp + all_goals simp only [Multiset.mem_toFinset, Finset.mem_union] + · rintro p (hpa | hpb) <;> apply prime_of_normalized_factor <;> assumption + · rintro p (hp | hp) q (hq | hq) hdvd <;> + rw [← normalize_normalized_factor _ hp, ← normalize_normalized_factor _ hq] <;> + exact + normalize_eq_normalize hdvd + ((prime_of_normalized_factor _ hp).irreducible.dvd_symm + (prime_of_normalized_factor _ hq).irreducible hdvd) + +end Multiplicative + +end UniqueFactorizationMonoid diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean new file mode 100644 index 0000000000000..dbc52cb56fad1 --- /dev/null +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean @@ -0,0 +1,126 @@ +/- +Copyright (c) 2018 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Jens Wagemaker, Aaron Anderson +-/ +import Mathlib.RingTheory.Multiplicity +import Mathlib.RingTheory.UniqueFactorizationDomain.NormalizedFactors + +/-! +# Unique factorization and multiplicity + +## Main results + +* `UniqueFactorizationMonoid.emultiplicity_eq_count_normalizedFactors`: The multiplicity of an + irreducible factor of a nonzero element is exactly the number of times the normalized factor + occurs in the `normalizedFactors`. +-/ + + +variable {α : Type*} + +local infixl:50 " ~ᵤ " => Associated + +theorem WfDvdMonoid.max_power_factor' [CommMonoidWithZero α] [WfDvdMonoid α] {a₀ x : α} + (h : a₀ ≠ 0) (hx : ¬IsUnit x) : ∃ (n : ℕ) (a : α), ¬x ∣ a ∧ a₀ = x ^ n * a := by + obtain ⟨a, ⟨n, rfl⟩, hm⟩ := wellFounded_dvdNotUnit.has_min + {a | ∃ n, x ^ n * a = a₀} ⟨a₀, 0, by rw [pow_zero, one_mul]⟩ + refine ⟨n, a, ?_, rfl⟩; rintro ⟨d, rfl⟩ + exact hm d ⟨n + 1, by rw [pow_succ, mul_assoc]⟩ + ⟨(right_ne_zero_of_mul <| right_ne_zero_of_mul h), x, hx, mul_comm _ _⟩ + +theorem WfDvdMonoid.max_power_factor [CommMonoidWithZero α] [WfDvdMonoid α] {a₀ x : α} + (h : a₀ ≠ 0) (hx : Irreducible x) : ∃ (n : ℕ) (a : α), ¬x ∣ a ∧ a₀ = x ^ n * a := + max_power_factor' h hx.not_unit + +theorem multiplicity.finite_of_not_isUnit [CancelCommMonoidWithZero α] [WfDvdMonoid α] + {a b : α} (ha : ¬IsUnit a) (hb : b ≠ 0) : multiplicity.Finite a b := by + obtain ⟨n, c, ndvd, rfl⟩ := WfDvdMonoid.max_power_factor' hb ha + exact ⟨n, by rwa [pow_succ, mul_dvd_mul_iff_left (left_ne_zero_of_mul hb)]⟩ + +namespace UniqueFactorizationMonoid + +variable {R : Type*} [CancelCommMonoidWithZero R] [UniqueFactorizationMonoid R] + +section multiplicity + +variable [NormalizationMonoid R] + +open Multiset + +section + +theorem le_emultiplicity_iff_replicate_le_normalizedFactors {a b : R} {n : ℕ} (ha : Irreducible a) + (hb : b ≠ 0) : + ↑n ≤ emultiplicity a b ↔ replicate n (normalize a) ≤ normalizedFactors b := by + rw [← pow_dvd_iff_le_emultiplicity] + revert b + induction' n with n ih; · simp + intro b hb + constructor + · rintro ⟨c, rfl⟩ + rw [Ne, pow_succ', mul_assoc, mul_eq_zero, not_or] at hb + rw [pow_succ', mul_assoc, normalizedFactors_mul hb.1 hb.2, replicate_succ, + normalizedFactors_irreducible ha, singleton_add, cons_le_cons_iff, ← ih hb.2] + apply Dvd.intro _ rfl + · rw [Multiset.le_iff_exists_add] + rintro ⟨u, hu⟩ + rw [← (normalizedFactors_prod hb).dvd_iff_dvd_right, hu, prod_add, prod_replicate] + exact (Associated.pow_pow <| associated_normalize a).dvd.trans (Dvd.intro u.prod rfl) + +/-- The multiplicity of an irreducible factor of a nonzero element is exactly the number of times +the normalized factor occurs in the `normalizedFactors`. + +See also `count_normalizedFactors_eq` which expands the definition of `multiplicity` +to produce a specification for `count (normalizedFactors _) _`.. +-/ +theorem emultiplicity_eq_count_normalizedFactors [DecidableEq R] {a b : R} (ha : Irreducible a) + (hb : b ≠ 0) : emultiplicity a b = (normalizedFactors b).count (normalize a) := by + apply le_antisymm + · apply Order.le_of_lt_add_one + rw [← Nat.cast_one, ← Nat.cast_add, lt_iff_not_ge, ge_iff_le, + le_emultiplicity_iff_replicate_le_normalizedFactors ha hb, ← le_count_iff_replicate_le] + simp + rw [le_emultiplicity_iff_replicate_le_normalizedFactors ha hb, ← le_count_iff_replicate_le] + +end + +/-- The number of times an irreducible factor `p` appears in `normalizedFactors x` is defined by +the number of times it divides `x`. + +See also `multiplicity_eq_count_normalizedFactors` if `n` is given by `multiplicity p x`. +-/ +theorem count_normalizedFactors_eq [DecidableEq R] {p x : R} (hp : Irreducible p) + (hnorm : normalize p = p) {n : ℕ} (hle : p ^ n ∣ x) (hlt : ¬p ^ (n + 1) ∣ x) : + (normalizedFactors x).count p = n := by classical + by_cases hx0 : x = 0 + · simp [hx0] at hlt + apply Nat.cast_injective (R := ℕ∞) + convert (emultiplicity_eq_count_normalizedFactors hp hx0).symm + · exact hnorm.symm + exact (emultiplicity_eq_coe.mpr ⟨hle, hlt⟩).symm + +/-- The number of times an irreducible factor `p` appears in `normalizedFactors x` is defined by +the number of times it divides `x`. This is a slightly more general version of +`UniqueFactorizationMonoid.count_normalizedFactors_eq` that allows `p = 0`. + +See also `multiplicity_eq_count_normalizedFactors` if `n` is given by `multiplicity p x`. +-/ +theorem count_normalizedFactors_eq' [DecidableEq R] {p x : R} (hp : p = 0 ∨ Irreducible p) + (hnorm : normalize p = p) {n : ℕ} (hle : p ^ n ∣ x) (hlt : ¬p ^ (n + 1) ∣ x) : + (normalizedFactors x).count p = n := by + rcases hp with (rfl | hp) + · cases n + · exact count_eq_zero.2 (zero_not_mem_normalizedFactors _) + · rw [zero_pow (Nat.succ_ne_zero _)] at hle hlt + exact absurd hle hlt + · exact count_normalizedFactors_eq hp hnorm hle hlt + +end multiplicity + +/-- Deprecated. Use `WfDvdMonoid.max_power_factor` instead. -/ +@[deprecated WfDvdMonoid.max_power_factor (since := "2024-03-01")] +theorem max_power_factor {a₀ x : R} (h : a₀ ≠ 0) (hx : Irreducible x) : + ∃ n : ℕ, ∃ a : R, ¬x ∣ a ∧ a₀ = x ^ n * a := WfDvdMonoid.max_power_factor h hx + +end UniqueFactorizationMonoid diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean new file mode 100644 index 0000000000000..680aa8f4249fe --- /dev/null +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean @@ -0,0 +1,62 @@ +/- +Copyright (c) 2018 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Jens Wagemaker, Aaron Anderson +-/ +import Mathlib.Data.ENat.Basic +import Mathlib.Data.Nat.Factors +import Mathlib.RingTheory.UniqueFactorizationDomain.NormalizedFactors + +/-! +# Unique factorization of natural numbers + +## Main definitions + +* `Nat.instUniqueFactorizationMonoid`: the natural numbers have unique factorization +-/ + +variable {α : Type*} + +namespace Nat + +instance instWfDvdMonoid : WfDvdMonoid ℕ where + wf := by + refine RelHomClass.wellFounded + (⟨fun x : ℕ => if x = 0 then (⊤ : ℕ∞) else x, ?_⟩ : DvdNotUnit →r (· < ·)) wellFounded_lt + intro a b h + cases' a with a + · exfalso + revert h + simp [DvdNotUnit] + cases b + · simpa [succ_ne_zero] using WithTop.coe_lt_top (a + 1) + cases' dvd_and_not_dvd_iff.2 h with h1 h2 + simp only [succ_ne_zero, cast_lt, if_false] + refine lt_of_le_of_ne (Nat.le_of_dvd (Nat.succ_pos _) h1) fun con => h2 ?_ + rw [con] + +instance instUniqueFactorizationMonoid : UniqueFactorizationMonoid ℕ where + irreducible_iff_prime := Nat.irreducible_iff_prime + +open UniqueFactorizationMonoid + +lemma factors_eq : ∀ n : ℕ, normalizedFactors n = n.primeFactorsList + | 0 => by simp + | n + 1 => by + rw [← Multiset.rel_eq, ← associated_eq_eq] + apply UniqueFactorizationMonoid.factors_unique irreducible_of_normalized_factor _ + · rw [Multiset.prod_coe, Nat.prod_primeFactorsList n.succ_ne_zero] + exact normalizedFactors_prod n.succ_ne_zero + · intro x hx + rw [Nat.irreducible_iff_prime, ← Nat.prime_iff] + exact Nat.prime_of_mem_primeFactorsList hx + +lemma factors_multiset_prod_of_irreducible {s : Multiset ℕ} (h : ∀ x : ℕ, x ∈ s → Irreducible x) : + normalizedFactors s.prod = s := by + rw [← Multiset.rel_eq, ← associated_eq_eq] + apply UniqueFactorizationMonoid.factors_unique irreducible_of_normalized_factor h + (normalizedFactors_prod _) + rw [Ne, Multiset.prod_eq_zero_iff] + exact fun con ↦ not_irreducible_zero (h 0 con) + +end Nat diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/NormalizedFactors.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/NormalizedFactors.lean new file mode 100644 index 0000000000000..3d30df0c3b813 --- /dev/null +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/NormalizedFactors.lean @@ -0,0 +1,329 @@ +/- +Copyright (c) 2018 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Jens Wagemaker, Aaron Anderson +-/ +import Mathlib.Algebra.GCDMonoid.Basic +import Mathlib.RingTheory.UniqueFactorizationDomain.Basic + +/-! +# Unique factorization and normalization + +## Main definitions +* `UniqueFactorizationMonoid.normalizedFactors`: choose a multiset of prime factors that are unique + by normalizing them. +* `UniqueFactorizationMonoid.normalizationMonoid`: choose a way of normalizing the elements of a UFM +-/ + + +variable {α : Type*} + +local infixl:50 " ~ᵤ " => Associated + +namespace UniqueFactorizationMonoid + +variable [CancelCommMonoidWithZero α] [NormalizationMonoid α] +variable [UniqueFactorizationMonoid α] + +/-- Noncomputably determines the multiset of prime factors. -/ +noncomputable def normalizedFactors (a : α) : Multiset α := + Multiset.map normalize <| factors a + +/-- An arbitrary choice of factors of `x : M` is exactly the (unique) normalized set of factors, +if `M` has a trivial group of units. -/ +@[simp] +theorem factors_eq_normalizedFactors {M : Type*} [CancelCommMonoidWithZero M] + [UniqueFactorizationMonoid M] [Subsingleton Mˣ] (x : M) : factors x = normalizedFactors x := by + unfold normalizedFactors + convert (Multiset.map_id (factors x)).symm + ext p + exact normalize_eq p + +theorem normalizedFactors_prod {a : α} (ane0 : a ≠ 0) : + Associated (normalizedFactors a).prod a := by + rw [normalizedFactors, factors, dif_neg ane0] + refine Associated.trans ?_ (Classical.choose_spec (exists_prime_factors a ane0)).2 + rw [← Associates.mk_eq_mk_iff_associated, ← Associates.prod_mk, ← Associates.prod_mk, + Multiset.map_map] + congr 2 + ext + rw [Function.comp_apply, Associates.mk_normalize] + +theorem prime_of_normalized_factor {a : α} : ∀ x : α, x ∈ normalizedFactors a → Prime x := by + rw [normalizedFactors, factors] + split_ifs with ane0; · simp + intro x hx; rcases Multiset.mem_map.1 hx with ⟨y, ⟨hy, rfl⟩⟩ + rw [(normalize_associated _).prime_iff] + exact (Classical.choose_spec (UniqueFactorizationMonoid.exists_prime_factors a ane0)).1 y hy + +theorem irreducible_of_normalized_factor {a : α} : + ∀ x : α, x ∈ normalizedFactors a → Irreducible x := fun x h => + (prime_of_normalized_factor x h).irreducible + +theorem normalize_normalized_factor {a : α} : + ∀ x : α, x ∈ normalizedFactors a → normalize x = x := by + rw [normalizedFactors, factors] + split_ifs with h; · simp + intro x hx + obtain ⟨y, _, rfl⟩ := Multiset.mem_map.1 hx + apply normalize_idem + +theorem normalizedFactors_irreducible {a : α} (ha : Irreducible a) : + normalizedFactors a = {normalize a} := by + obtain ⟨p, a_assoc, hp⟩ := + prime_factors_irreducible ha ⟨prime_of_normalized_factor, normalizedFactors_prod ha.ne_zero⟩ + have p_mem : p ∈ normalizedFactors a := by + rw [hp] + exact Multiset.mem_singleton_self _ + convert hp + rwa [← normalize_normalized_factor p p_mem, normalize_eq_normalize_iff, dvd_dvd_iff_associated] + +theorem normalizedFactors_eq_of_dvd (a : α) : + ∀ᵉ (p ∈ normalizedFactors a) (q ∈ normalizedFactors a), p ∣ q → p = q := by + intro p hp q hq hdvd + convert normalize_eq_normalize hdvd + ((prime_of_normalized_factor _ hp).irreducible.dvd_symm + (prime_of_normalized_factor _ hq).irreducible hdvd) <;> + apply (normalize_normalized_factor _ ‹_›).symm + +theorem exists_mem_normalizedFactors_of_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) : + p ∣ a → ∃ q ∈ normalizedFactors a, p ~ᵤ q := fun ⟨b, hb⟩ => + have hb0 : b ≠ 0 := fun hb0 => by simp_all + have : Multiset.Rel Associated (p ::ₘ normalizedFactors b) (normalizedFactors a) := + factors_unique + (fun _ hx => + (Multiset.mem_cons.1 hx).elim (fun h => h.symm ▸ hp) (irreducible_of_normalized_factor _)) + irreducible_of_normalized_factor + (Associated.symm <| + calc + Multiset.prod (normalizedFactors a) ~ᵤ a := normalizedFactors_prod ha0 + _ = p * b := hb + _ ~ᵤ Multiset.prod (p ::ₘ normalizedFactors b) := by + rw [Multiset.prod_cons] + exact (normalizedFactors_prod hb0).symm.mul_left _ + ) + Multiset.exists_mem_of_rel_of_mem this (by simp) + +theorem exists_mem_normalizedFactors {x : α} (hx : x ≠ 0) (h : ¬IsUnit x) : + ∃ p, p ∈ normalizedFactors x := by + obtain ⟨p', hp', hp'x⟩ := WfDvdMonoid.exists_irreducible_factor h hx + obtain ⟨p, hp, _⟩ := exists_mem_normalizedFactors_of_dvd hx hp' hp'x + exact ⟨p, hp⟩ + +@[simp] +theorem normalizedFactors_zero : normalizedFactors (0 : α) = 0 := by + simp [normalizedFactors, factors] + +@[simp] +theorem normalizedFactors_one : normalizedFactors (1 : α) = 0 := by + cases' subsingleton_or_nontrivial α with h h + · dsimp [normalizedFactors, factors] + simp [Subsingleton.elim (1 : α) 0] + · rw [← Multiset.rel_zero_right] + apply factors_unique irreducible_of_normalized_factor + · intro x hx + exfalso + apply Multiset.not_mem_zero x hx + · apply normalizedFactors_prod one_ne_zero + +@[simp] +theorem normalizedFactors_mul {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : + normalizedFactors (x * y) = normalizedFactors x + normalizedFactors y := by + have h : (normalize : α → α) = Associates.out ∘ Associates.mk := by + ext + rw [Function.comp_apply, Associates.out_mk] + rw [← Multiset.map_id' (normalizedFactors (x * y)), ← Multiset.map_id' (normalizedFactors x), ← + Multiset.map_id' (normalizedFactors y), ← Multiset.map_congr rfl normalize_normalized_factor, ← + Multiset.map_congr rfl normalize_normalized_factor, ← + Multiset.map_congr rfl normalize_normalized_factor, ← Multiset.map_add, h, ← + Multiset.map_map Associates.out, eq_comm, ← Multiset.map_map Associates.out] + refine congr rfl ?_ + apply Multiset.map_mk_eq_map_mk_of_rel + apply factors_unique + · intro x hx + rcases Multiset.mem_add.1 hx with (hx | hx) <;> exact irreducible_of_normalized_factor x hx + · exact irreducible_of_normalized_factor + · rw [Multiset.prod_add] + exact + ((normalizedFactors_prod hx).mul_mul (normalizedFactors_prod hy)).trans + (normalizedFactors_prod (mul_ne_zero hx hy)).symm + +@[simp] +theorem normalizedFactors_pow {x : α} (n : ℕ) : + normalizedFactors (x ^ n) = n • normalizedFactors x := by + induction' n with n ih + · simp + by_cases h0 : x = 0 + · simp [h0, zero_pow n.succ_ne_zero, smul_zero] + rw [pow_succ', succ_nsmul', normalizedFactors_mul h0 (pow_ne_zero _ h0), ih] + +theorem _root_.Irreducible.normalizedFactors_pow {p : α} (hp : Irreducible p) (k : ℕ) : + normalizedFactors (p ^ k) = Multiset.replicate k (normalize p) := by + rw [UniqueFactorizationMonoid.normalizedFactors_pow, normalizedFactors_irreducible hp, + Multiset.nsmul_singleton] + +theorem normalizedFactors_prod_eq (s : Multiset α) (hs : ∀ a ∈ s, Irreducible a) : + normalizedFactors s.prod = s.map normalize := by + induction' s using Multiset.induction with a s ih + · rw [Multiset.prod_zero, normalizedFactors_one, Multiset.map_zero] + · have ia := hs a (Multiset.mem_cons_self a _) + have ib := fun b h => hs b (Multiset.mem_cons_of_mem h) + obtain rfl | ⟨b, hb⟩ := s.empty_or_exists_mem + · rw [Multiset.cons_zero, Multiset.prod_singleton, Multiset.map_singleton, + normalizedFactors_irreducible ia] + haveI := nontrivial_of_ne b 0 (ib b hb).ne_zero + rw [Multiset.prod_cons, Multiset.map_cons, + normalizedFactors_mul ia.ne_zero (Multiset.prod_ne_zero fun h => (ib 0 h).ne_zero rfl), + normalizedFactors_irreducible ia, ih ib, Multiset.singleton_add] + +theorem dvd_iff_normalizedFactors_le_normalizedFactors {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : + x ∣ y ↔ normalizedFactors x ≤ normalizedFactors y := by + constructor + · rintro ⟨c, rfl⟩ + simp [hx, right_ne_zero_of_mul hy] + · rw [← (normalizedFactors_prod hx).dvd_iff_dvd_left, ← + (normalizedFactors_prod hy).dvd_iff_dvd_right] + apply Multiset.prod_dvd_prod_of_le + +theorem _root_.Associated.normalizedFactors_eq {a b : α} (h : Associated a b) : + normalizedFactors a = normalizedFactors b := by + unfold normalizedFactors + have h' : ⇑(normalize (α := α)) = Associates.out ∘ Associates.mk := funext Associates.out_mk + rw [h', ← Multiset.map_map, ← Multiset.map_map, + Associates.rel_associated_iff_map_eq_map.mp (factors_rel_of_associated h)] + +theorem associated_iff_normalizedFactors_eq_normalizedFactors {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : + x ~ᵤ y ↔ normalizedFactors x = normalizedFactors y := + ⟨Associated.normalizedFactors_eq, fun h => + (normalizedFactors_prod hx).symm.trans (_root_.trans (by rw [h]) (normalizedFactors_prod hy))⟩ + +theorem normalizedFactors_of_irreducible_pow {p : α} (hp : Irreducible p) (k : ℕ) : + normalizedFactors (p ^ k) = Multiset.replicate k (normalize p) := by + rw [normalizedFactors_pow, normalizedFactors_irreducible hp, Multiset.nsmul_singleton] + +theorem zero_not_mem_normalizedFactors (x : α) : (0 : α) ∉ normalizedFactors x := fun h => + Prime.ne_zero (prime_of_normalized_factor _ h) rfl + +theorem dvd_of_mem_normalizedFactors {a p : α} (H : p ∈ normalizedFactors a) : p ∣ a := by + by_cases hcases : a = 0 + · rw [hcases] + exact dvd_zero p + · exact dvd_trans (Multiset.dvd_prod H) (Associated.dvd (normalizedFactors_prod hcases)) + +theorem mem_normalizedFactors_iff [Subsingleton αˣ] {p x : α} (hx : x ≠ 0) : + p ∈ normalizedFactors x ↔ Prime p ∧ p ∣ x := by + constructor + · intro h + exact ⟨prime_of_normalized_factor p h, dvd_of_mem_normalizedFactors h⟩ + · rintro ⟨hprime, hdvd⟩ + obtain ⟨q, hqmem, hqeq⟩ := exists_mem_normalizedFactors_of_dvd hx hprime.irreducible hdvd + rw [associated_iff_eq] at hqeq + exact hqeq ▸ hqmem + +theorem exists_associated_prime_pow_of_unique_normalized_factor {p r : α} + (h : ∀ {m}, m ∈ normalizedFactors r → m = p) (hr : r ≠ 0) : ∃ i : ℕ, Associated (p ^ i) r := by + use Multiset.card.toFun (normalizedFactors r) + have := UniqueFactorizationMonoid.normalizedFactors_prod hr + rwa [Multiset.eq_replicate_of_mem fun b => h, Multiset.prod_replicate] at this + +theorem normalizedFactors_prod_of_prime [Subsingleton αˣ] {m : Multiset α} + (h : ∀ p ∈ m, Prime p) : normalizedFactors m.prod = m := by + cases subsingleton_or_nontrivial α + · obtain rfl : m = 0 := by + refine Multiset.eq_zero_of_forall_not_mem fun x hx ↦ ?_ + simpa [Subsingleton.elim x 0] using h x hx + simp + · simpa only [← Multiset.rel_eq, ← associated_eq_eq] using + prime_factors_unique prime_of_normalized_factor h + (normalizedFactors_prod (m.prod_ne_zero_of_prime h)) + +theorem mem_normalizedFactors_eq_of_associated {a b c : α} (ha : a ∈ normalizedFactors c) + (hb : b ∈ normalizedFactors c) (h : Associated a b) : a = b := by + rw [← normalize_normalized_factor a ha, ← normalize_normalized_factor b hb, + normalize_eq_normalize_iff] + exact Associated.dvd_dvd h + +@[simp] +theorem normalizedFactors_pos (x : α) (hx : x ≠ 0) : 0 < normalizedFactors x ↔ ¬IsUnit x := by + constructor + · intro h hx + obtain ⟨p, hp⟩ := Multiset.exists_mem_of_ne_zero h.ne' + exact + (prime_of_normalized_factor _ hp).not_unit + (isUnit_of_dvd_unit (dvd_of_mem_normalizedFactors hp) hx) + · intro h + obtain ⟨p, hp⟩ := exists_mem_normalizedFactors hx h + exact + bot_lt_iff_ne_bot.mpr + (mt Multiset.eq_zero_iff_forall_not_mem.mp (not_forall.mpr ⟨p, not_not.mpr hp⟩)) + +theorem dvdNotUnit_iff_normalizedFactors_lt_normalizedFactors {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : + DvdNotUnit x y ↔ normalizedFactors x < normalizedFactors y := by + constructor + · rintro ⟨_, c, hc, rfl⟩ + simp only [hx, right_ne_zero_of_mul hy, normalizedFactors_mul, Ne, not_false_iff, + lt_add_iff_pos_right, normalizedFactors_pos, hc] + · intro h + exact + dvdNotUnit_of_dvd_of_not_dvd + ((dvd_iff_normalizedFactors_le_normalizedFactors hx hy).mpr h.le) + (mt (dvd_iff_normalizedFactors_le_normalizedFactors hy hx).mp h.not_le) + +theorem normalizedFactors_multiset_prod (s : Multiset α) (hs : 0 ∉ s) : + normalizedFactors (s.prod) = (s.map normalizedFactors).sum := by + cases subsingleton_or_nontrivial α + · obtain rfl : s = 0 := by + apply Multiset.eq_zero_of_forall_not_mem + intro _ + convert hs + simp + induction s using Multiset.induction with + | empty => simp + | cons _ _ IH => + rw [Multiset.prod_cons, Multiset.map_cons, Multiset.sum_cons, normalizedFactors_mul, IH] + · exact fun h ↦ hs (Multiset.mem_cons_of_mem h) + · exact fun h ↦ hs (h ▸ Multiset.mem_cons_self _ _) + · apply Multiset.prod_ne_zero + exact fun h ↦ hs (Multiset.mem_cons_of_mem h) + +end UniqueFactorizationMonoid + +namespace UniqueFactorizationMonoid + +open scoped Classical + +open Multiset Associates + +variable [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] + +/-- Noncomputably defines a `normalizationMonoid` structure on a `UniqueFactorizationMonoid`. -/ +protected noncomputable def normalizationMonoid : NormalizationMonoid α := + normalizationMonoidOfMonoidHomRightInverse + { toFun := fun a : Associates α => + if a = 0 then 0 + else + ((normalizedFactors a).map + (Classical.choose mk_surjective.hasRightInverse : Associates α → α)).prod + map_one' := by nontriviality α; simp + map_mul' := fun x y => by + by_cases hx : x = 0 + · simp [hx] + by_cases hy : y = 0 + · simp [hy] + simp [hx, hy] } + (by + intro x + dsimp + by_cases hx : x = 0 + · simp [hx] + have h : Associates.mkMonoidHom ∘ Classical.choose mk_surjective.hasRightInverse = + (id : Associates α → Associates α) := by + ext x + rw [Function.comp_apply, mkMonoidHom_apply, + Classical.choose_spec mk_surjective.hasRightInverse x] + rfl + rw [if_neg hx, ← mkMonoidHom_apply, MonoidHom.map_multiset_prod, map_map, h, map_id, ← + associated_iff_eq] + apply normalizedFactors_prod hx) + +end UniqueFactorizationMonoid diff --git a/MathlibTest/Variable.lean b/MathlibTest/Variable.lean index 327d36efd228e..e93a74e81fbec 100644 --- a/MathlibTest/Variable.lean +++ b/MathlibTest/Variable.lean @@ -1,8 +1,10 @@ -import Mathlib.Tactic.Variable -import Mathlib.Algebra.Module.Defs import Mathlib.Algebra.Algebra.Defs +import Mathlib.Algebra.Field.Defs +import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.Ring.Regular import Mathlib.Algebra.Module.LinearMap.Basic -import Mathlib.RingTheory.UniqueFactorizationDomain +import Mathlib.RingTheory.UniqueFactorizationDomain.Basic +import Mathlib.Tactic.Variable set_option autoImplicit true namespace Tests From 9da26514ae8782fb1fcf86b70fb4acd92c21c7ba Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 20 Nov 2024 11:43:02 +0000 Subject: [PATCH 142/829] chore: bump dependencies (#19279) --- Mathlib/Algebra/Order/Group/Unbundled/Basic.lean | 1 + lake-manifest.json | 2 +- 2 files changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean index 03c9d7579e5f1..5b90caf6e8964 100644 --- a/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean @@ -537,6 +537,7 @@ section LE variable [LE α] [MulLeftMono α] {a b c d : α} +set_option linter.docPrime false in @[to_additive sub_le_sub_iff] theorem div_le_div_iff' : a / b ≤ c / d ↔ a * d ≤ c * b := by simpa only [div_eq_mul_inv] using mul_inv_le_mul_inv_iff' diff --git a/lake-manifest.json b/lake-manifest.json index cfc65dc4caaa9..d40bdeb3be7db 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "86d0d0584f5cd165353e2f8a30c455cd0e168ac2", + "rev": "d7caecce0d0f003fd5e9cce9a61f1dd6ba83142b", "name": "LeanSearchClient", "manifestFile": "lake-manifest.json", "inputRev": "main", From cd4d2185e67b7387e303150dab86521007c56fb6 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Wed, 20 Nov 2024 12:32:28 +0000 Subject: [PATCH 143/829] feat(KrullDimension): concrete calculations (#19210) Part of https://github.com/leanprover-community/mathlib4/pull/15524. This adds concrete calculations of the height and Krull Dimension for Nat, Int, WithTop, WithBot and ENat. From the Carleson project. --- Mathlib/Order/KrullDimension.lean | 158 +++++++++++++++++++++++++++++- 1 file changed, 156 insertions(+), 2 deletions(-) diff --git a/Mathlib/Order/KrullDimension.lean b/Mathlib/Order/KrullDimension.lean index 836303c212832..249a8d2cc0301 100644 --- a/Mathlib/Order/KrullDimension.lean +++ b/Mathlib/Order/KrullDimension.lean @@ -4,9 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang, Fangming Li, Joachim Breitner -/ -import Mathlib.Order.RelSeries -import Mathlib.Order.Minimal +import Mathlib.Algebra.Order.Group.Int import Mathlib.Data.ENat.Lattice +import Mathlib.Order.Minimal +import Mathlib.Order.RelSeries /-! # Krull dimension of a preordered set and height of an element @@ -648,4 +649,157 @@ lemma coheight_bot_eq_krullDim [OrderBot α] : coheight (⊥ : α) = krullDim α end krullDim +/-! +## Concrete calculations +-/ + +section calculations + +variable {α : Type*} [Preorder α] + +/- +These two lemmas could possibly be used to simplify the subsequent calculations, +especially once the `Set.encard` api is richer. + +(Commented out to avoid importing modules purely for `proof_wanted`.) +proof_wanted height_of_linearOrder {α : Type*} [LinearOrder α] (a : α) : + height a = (Set.Iio a).encard + +proof_wanted coheight_of_linearOrder {α : Type*} [LinearOrder α] (a : α) : + coheight a = (Set.Ioi a).encard +-/ + +@[simp] lemma height_nat (n : ℕ) : height n = n := by + induction n using Nat.strongRecOn with | ind n ih => + apply le_antisymm + · apply height_le_coe_iff.mpr + simp +contextual only [ih, Nat.cast_lt, implies_true] + · exact length_le_height_last (p := LTSeries.range n) + +@[simp] lemma coheight_of_noMaxOrder [NoMaxOrder α] (a : α) : coheight a = ⊤ := by + obtain ⟨f, hstrictmono⟩ := Nat.exists_strictMono ↑(Set.Ioi a) + apply coheight_eq_top_iff.mpr + intro m + use {length := m, toFun := fun i => if i = 0 then a else f i, step := ?step } + case h => simp [RelSeries.head] + case step => + intro ⟨i, hi⟩ + by_cases hzero : i = 0 + · subst i + exact (f 1).prop + · suffices f i < f (i + 1) by simp [Fin.ext_iff, hzero, this] + apply hstrictmono + omega + +@[simp] lemma height_of_noMinOrder [NoMinOrder α] (a : α) : height a = ⊤ := + -- Implementation note: Here it's a bit easier to define the coheight variant first + coheight_of_noMaxOrder (α := αᵒᵈ) a + +@[simp] lemma krullDim_of_noMaxOrder [Nonempty α] [NoMaxOrder α] : krullDim α = ⊤ := by + simp [krullDim_eq_iSup_coheight, coheight_of_noMaxOrder] + +@[simp] lemma krullDim_of_noMinOrder [Nonempty α] [NoMinOrder α] : krullDim α = ⊤ := by + simp [krullDim_eq_iSup_height, height_of_noMinOrder] + +lemma coheight_nat (n : ℕ) : coheight n = ⊤ := coheight_of_noMaxOrder .. + +lemma krullDim_nat : krullDim ℕ = ⊤ := krullDim_of_noMaxOrder .. + +lemma height_int (n : ℤ) : height n = ⊤ := height_of_noMinOrder .. + +lemma coheight_int (n : ℤ) : coheight n = ⊤ := coheight_of_noMaxOrder .. + +lemma krullDim_int : krullDim ℤ = ⊤ := krullDim_of_noMaxOrder .. + +@[simp] lemma height_coe_withBot (x : α) : height (x : WithBot α) = height x + 1 := by + apply le_antisymm + · apply height_le + intro p hlast + wlog hlenpos : p.length ≠ 0 + · simp_all + -- essentially p' := (p.drop 1).map unbot + let p' : LTSeries α := { + length := p.length - 1 + toFun := fun ⟨i, hi⟩ => (p ⟨i+1, by omega⟩).unbot (by + apply LT.lt.ne_bot (a := p.head) + apply p.strictMono + exact compare_gt_iff_gt.mp rfl) + step := fun i => by simpa [WithBot.unbot_lt_iff] using p.step ⟨i + 1, by omega⟩ } + have hlast' : p'.last = x := by + simp only [RelSeries.last, Fin.val_last, WithBot.unbot_eq_iff, ← hlast, Fin.last] + congr + omega + suffices p'.length ≤ height p'.last by + simpa [p', hlast'] using this + apply length_le_height_last + · rw [height_add_const] + apply iSup₂_le + intro p hlast + let p' := (p.map _ WithBot.coe_strictMono).cons ⊥ (by simp) + apply le_iSup₂_of_le p' (by simp [p', hlast]) (by simp [p']) + +@[simp] lemma coheight_coe_withTop (x : α) : coheight (x : WithTop α) = coheight x + 1 := + height_coe_withBot (α := αᵒᵈ) x + +@[simp] lemma height_coe_withTop (x : α) : height (x : WithTop α) = height x := by + apply le_antisymm + · apply height_le + intro p hlast + -- essentially p' := p.map untop + let p' : LTSeries α := { + length := p.length + toFun := fun i => (p i).untop (by + apply WithTop.lt_top_iff_ne_top.mp + apply lt_of_le_of_lt + · exact p.monotone (Fin.le_last _) + · rw [RelSeries.last] at hlast + simp [hlast]) + step := fun i => by simpa only [WithTop.untop_lt_iff, WithTop.coe_untop] using p.step i } + have hlast' : p'.last = x := by + simp only [RelSeries.last, Fin.val_last, WithTop.untop_eq_iff, ← hlast] + suffices p'.length ≤ height p'.last by + rw [hlast'] at this + simpa [p'] using this + apply length_le_height_last + · apply height_le + intro p hlast + let p' := p.map _ WithTop.coe_strictMono + apply le_iSup₂_of_le p' (by simp [p', hlast]) (by simp [p']) + +@[simp] lemma coheight_coe_withBot (x : α) : coheight (x : WithBot α) = coheight x := + height_coe_withTop (α := αᵒᵈ) x + +@[simp] lemma krullDim_WithTop [Nonempty α] : krullDim (WithTop α) = krullDim α + 1 := by + rw [← height_top_eq_krullDim, krullDim_eq_iSup_height_of_nonempty, height_eq_iSup_lt_height] + norm_cast + simp_rw [WithTop.lt_top_iff_ne_top] + rw [ENat.iSup_add, iSup_subtype'] + symm + apply Equiv.withTopSubtypeNe.symm.iSup_congr + simp + +@[simp] lemma krullDim_withBot [Nonempty α] : krullDim (WithBot α) = krullDim α + 1 := by + conv_lhs => rw [← krullDim_orderDual] + conv_rhs => rw [← krullDim_orderDual] + exact krullDim_WithTop (α := αᵒᵈ) + +@[simp] +lemma krullDim_enat : krullDim ℕ∞ = ⊤ := by + show (krullDim (WithTop ℕ) = ⊤) + simp only [krullDim_WithTop, krullDim_nat] + rfl + +@[simp] +lemma height_enat (n : ℕ∞) : height n = n := by + cases n with + | top => simp only [← WithBot.coe_eq_coe, height_top_eq_krullDim, krullDim_enat, WithBot.coe_top] + | coe n => exact (height_coe_withTop _).trans (height_nat _) + +@[simp] +lemma coheight_coe_enat (n : ℕ) : coheight (n : ℕ∞) = ⊤ := by + apply (coheight_coe_withTop _).trans + simp only [Nat.cast_id, coheight_nat, top_add] + +end calculations + end Order From 79ee32ececd77acd6e5c7ef387de3771c9790787 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Wed, 20 Nov 2024 12:45:10 +0000 Subject: [PATCH 144/829] feat: cardinality of `Field.Emb` for algebraic extensions of infinite separable degree (#9480) We show that the number of embeddings of an algebraic extension E/F with infinite `sepDegree F E` is equal to `2 ^ sepDegree F E`. Since `separableClosure F E` has the same number of F-embeddings as E, we may assume E/F is separable, so `sepDegree F E = Module.rank F E` ([E:F]). The first part of the proof (up to `two_le_deg`) shows the existence of a transfinite filtration of E/F by subextensions of length equal to [E:F], such that each step is finite and nontrivial (has degree at least 2). The second part (up to `embEquivPi`) shows the embeddings of these subextensions form an inverse system satisfying the conditions of `InverseSystem.globalEquiv`, and therefore the number of embeddings of E/F is equal to the product of the degrees of the steps. We conclude the file with five theorems stating the main result for separable and general algebraic extensions, incorporating the finite case or not. [Zulip](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/Separability.20degree.3F/near/411486335) Co-authored-by: Junyan Xu --- Mathlib.lean | 1 + Mathlib/FieldTheory/CardinalEmb.lean | 342 +++++++++++++++++++++++ Mathlib/FieldTheory/SeparableDegree.lean | 25 +- Mathlib/LinearAlgebra/FreeAlgebra.lean | 12 + Mathlib/Order/SuccPred/Basic.lean | 4 +- 5 files changed, 367 insertions(+), 17 deletions(-) create mode 100644 Mathlib/FieldTheory/CardinalEmb.lean diff --git a/Mathlib.lean b/Mathlib.lean index 0426236be0a65..db0e4160d4515 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2916,6 +2916,7 @@ import Mathlib.FieldTheory.AbsoluteGaloisGroup import Mathlib.FieldTheory.Adjoin import Mathlib.FieldTheory.AlgebraicClosure import Mathlib.FieldTheory.AxGrothendieck +import Mathlib.FieldTheory.CardinalEmb import Mathlib.FieldTheory.Cardinality import Mathlib.FieldTheory.ChevalleyWarning import Mathlib.FieldTheory.Differential.Basic diff --git a/Mathlib/FieldTheory/CardinalEmb.lean b/Mathlib/FieldTheory/CardinalEmb.lean new file mode 100644 index 0000000000000..1d4c21d21c883 --- /dev/null +++ b/Mathlib/FieldTheory/CardinalEmb.lean @@ -0,0 +1,342 @@ +/- +Copyright (c) 2024 Junyan Xu. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Junyan Xu +-/ +import Mathlib.FieldTheory.SeparableClosure +import Mathlib.FieldTheory.PurelyInseparable +import Mathlib.LinearAlgebra.FreeAlgebra +import Mathlib.Order.Interval.Set.WithBotTop +import Mathlib.Order.DirectedInverseSystem + +/-! +# Number of embeddings of an algebraic extension of infinite separable degree + +## Main results + +- `Field.Emb.cardinal_eq_two_pow_rank` : if `E/F` is an algebraic separable field extension +of infinite degree, then `#(Field.Emb F E) = 2 ^ Module.rank F E`. +This is in contrast to the case of finite degree, where `#(Field.Emb F E) = Module.rank F E`. + +- `Field.Emb.cardinal_eq_two_pow_sepDegree`: more generally, if `E/F` is an algebraic +extension of infinite separable degree, then `#(Field.Emb F E) = 2 ^ Field.sepDegree F E`. + +## Sketch of the proof + +We use a transfinite recursive construction that is fairly standard in set theory, but the author +has not seen similar arguments elsewhere in mathlib, and some parts proved tricky to formalize. + +The extension `E/F` can be filtered by intermediate fields indexed by a well-order: +simply put a well-order on a basis of `E/F`, and at each step, take the smallest basis element +that is not contained in the intermediate field generated by all previous elements, so that they +generate a strictly larger intermediate field together. This process can extend all the way +to the initial ordinal `ι` of the cardinal `Module.rank F E`, because the dimension of the +subalgebra generated by an infinite set cannot be greater than the cardinality of the set, and +in an algebraic extension, any subalgebra is a field. This is proven as `Algebra.rank_adjoin_le` +and used to show `leastExt` is a total function from `ι` to itself. It is probably +mathematically the most nontrivial part of the whole argument, but turned out easy to +formalize and was done the earliest. + +Once we have the filtration `E⟮ (Module.rank F E).ord.toType + +private local instance : SuccOrder ι := SuccOrder.ofLinearWellFoundedLT ι +local notation i"⁺" => succ i -- Note: conflicts with `PosPart` notation + +/-- A basis of E/F indexed by the initial ordinal. -/ +def wellOrderedBasis : Basis ι F E := + (chooseBasis F E).reindex + (Cardinal.eq.mp <| (mk_ord_toType _).trans <| rank_eq_card_chooseBasisIndex F E).some.symm + +local notation "b" => wellOrderedBasis F E +local notation "Ē" => AlgebraicClosure E + +variable {F E} + +theorem adjoin_basis_eq_top : adjoin F (range b) = ⊤ := + toSubalgebra_injective <| Subalgebra.toSubmodule_injective <| top_unique <| + (Basis.span_eq b).ge.trans <| (Algebra.span_le_adjoin F _).trans <| algebra_adjoin_le_adjoin _ _ + +section Algebraic + +variable [rank_inf : Fact (ℵ₀ ≤ Module.rank F E)] + +lemma noMaxOrder_rank_toType : NoMaxOrder ι := Cardinal.noMaxOrder Fact.out +attribute [local instance] noMaxOrder_rank_toType + +open _root_.Algebra (IsAlgebraic) +variable [IsAlgebraic F E] + +variable (F E) in +/-- `leastExt i` is defined to be the smallest `k : ι` that generates a nontrival extension over +(i.e. does not lie in) the subalgebra (= intermediate field) generated by all previous +`leastExt j`, `j < i`. For cardinality reasons, such `k` always exist if `ι` is infinite. -/ +def leastExt : ι → ι := + wellFounded_lt.fix fun i ih ↦ + let s := range fun j : Iio i ↦ b (ih j j.2) + wellFounded_lt.min {k | b k ∉ adjoin F s} <| by + rw [← compl_setOf, nonempty_compl]; by_contra! + simp_rw [eq_univ_iff_forall, mem_setOf] at this + have := adjoin_le_iff.mpr (range_subset_iff.mpr this) + rw [adjoin_basis_eq_top, ← eq_top_iff] at this + apply_fun Module.rank F at this + refine ne_of_lt ?_ this + conv_rhs => rw [topEquiv.toLinearEquiv.rank_eq] + have := mk_Iio_ord_toType i + obtain eq | lt := rank_inf.out.eq_or_lt + · replace this := mk_lt_aleph0_iff.mp (this.trans_eq eq.symm) + have : FiniteDimensional F (adjoin F s) := + finiteDimensional_adjoin fun x _ ↦ (IsAlgebraic.isAlgebraic x).isIntegral + exact (Module.rank_lt_aleph0 _ _).trans_eq eq + · exact (Subalgebra.equivOfEq _ _ <| adjoin_algebraic_toSubalgebra + fun x _ ↦ IsAlgebraic.isAlgebraic x)|>.toLinearEquiv.rank_eq.trans_lt <| + (Algebra.rank_adjoin_le _).trans_lt (max_lt (mk_range_le.trans_lt this) lt) + +local notation "φ" => leastExt F E + +section +local notation "E⟮<"i"⟯" => adjoin F (b ∘ φ '' Iio i) + +theorem isLeast_leastExt (i : ι) : IsLeast {k | b k ∉ E⟮ Field.Emb (E⟮ filtration i + +variable (F E) in +/-- The functor on `WithTop ι` given by embeddings of `E⟮ rw [← mk_ord_toType (Module.rank F E), ← prod_const'] + exact prod_le_prod _ _ fun i ↦ (Emb.Cardinal.deg_lt_aleph0 _).le + · conv_lhs => rw [← mk_ord_toType (Module.rank F E), ← prod_const'] + exact prod_le_prod _ _ Emb.Cardinal.two_le_deg + +theorem cardinal_eq_of_isSeparable [Algebra.IsSeparable F E] : + #(Field.Emb F E) = (fun c ↦ if ℵ₀ ≤ c then 2 ^ c else c) (Module.rank F E) := by + dsimp only; split_ifs with h + · exact cardinal_eq_two_pow_rank h + rw [not_le, ← IsNoetherian.iff_rank_lt_aleph0] at h + rw [← Module.finrank_eq_rank, ← toNat_eq_iff Module.finrank_pos.ne', + ← Nat.card, ← finSepDegree, finSepDegree_eq_finrank_of_isSeparable] + +theorem cardinal_eq_two_pow_sepDegree [Algebra.IsAlgebraic F E] + (rank_inf : ℵ₀ ≤ sepDegree F E) : #(Field.Emb F E) = 2 ^ sepDegree F E := by + rw [← cardinal_separableClosure, cardinal_eq_two_pow_rank rank_inf] + rfl + +theorem cardinal_eq [Algebra.IsAlgebraic F E] : + #(Field.Emb F E) = (fun c ↦ if ℵ₀ ≤ c then 2 ^ c else c) (sepDegree F E) := by + rw [← cardinal_separableClosure, cardinal_eq_of_isSeparable]; rfl + +end Field.Emb diff --git a/Mathlib/FieldTheory/SeparableDegree.lean b/Mathlib/FieldTheory/SeparableDegree.lean index 8025c4eff3c76..c6a67370781a2 100644 --- a/Mathlib/FieldTheory/SeparableDegree.lean +++ b/Mathlib/FieldTheory/SeparableDegree.lean @@ -24,26 +24,21 @@ This file contains basics about the separable degree of a field extension. advantage that `Field.Emb F E` lies in the same universe as `E` rather than the maximum over `F` and `E`). Usually denoted by $\operatorname{Emb}_F(E)$ in textbooks. - **Remark:** if `E / F` is not algebraic, then this definition makes no mathematical sense, - and if it is infinite, then its cardinality doesn't behave as expected (namely, not equal to the - field extension degree of `separableClosure F E / F`). For example, if $F = \mathbb{Q}$ and - $E = \mathbb{Q}( \mu_{p^\infty} )$, then $\operatorname{Emb}_F (E)$ is in bijection with - $\operatorname{Gal}(E/F)$, which is isomorphic to - $\mathbb{Z}_p^\times$, which is uncountable, while $[E:F]$ is countable. - - **TODO:** prove or disprove that if `E / F` is algebraic and `Emb F E` is infinite, then - `Field.Emb F E` has cardinality `2 ^ Module.rank F (separableClosure F E)`. - -- `Field.finSepDegree F E`: the (finite) separable degree $[E:F]_s$ of an algebraic extension - `E / F` of fields, defined to be the number of `F`-algebra homomorphisms from `E` to the algebraic +- `Field.finSepDegree F E`: the (finite) separable degree $[E:F]_s$ of an extension `E / F` + of fields, defined to be the number of `F`-algebra homomorphisms from `E` to the algebraic closure of `E`, as a natural number. It is zero if `Field.Emb F E` is not finite. Note that if `E / F` is not algebraic, then this definition makes no mathematical sense. **Remark:** the `Cardinal`-valued, potentially infinite separable degree `Field.sepDegree F E` for a general algebraic extension `E / F` is defined to be the degree of `L / F`, where `L` is - the (relative) separable closure `separableClosure F E` of `F` in `E`, which is not defined in - this file yet. Later we will show that (`Field.finSepDegree_eq`), if `Field.Emb F E` is finite, - then these two definitions coincide. + the separable closure of `F` in `E`, which is not defined in this file yet. Later we + will show that (`Field.finSepDegree_eq`), if `Field.Emb F E` is finite, then these two + definitions coincide. If `E / F` is algebraic with infinite separable degree, we have + `#(Field.Emb F E) = 2 ^ Field.sepDegree F E` instead. + (See `Field.Emb.cardinal_eq_two_pow_sepDegree` in another file.) For example, if + $F = \mathbb{Q}$ and $E = \mathbb{Q}( \mu_{p^\infty} )$, then $\operatorname{Emb}_F (E)$ + is in bijection with $\operatorname{Gal}(E/F)$, which is isomorphic to + $\mathbb{Z}_p^\times$, which is uncountable, whereas $ [E:F] $ is countable. - `Polynomial.natSepDegree`: the separable degree of a polynomial is a natural number, defined to be the number of distinct roots of it over its splitting field. diff --git a/Mathlib/LinearAlgebra/FreeAlgebra.lean b/Mathlib/LinearAlgebra/FreeAlgebra.lean index 02bbf0606d764..727a91d91de6a 100644 --- a/Mathlib/LinearAlgebra/FreeAlgebra.lean +++ b/Mathlib/LinearAlgebra/FreeAlgebra.lean @@ -45,3 +45,15 @@ theorem rank_eq [CommRing R] [Nontrivial R] : Cardinal.lift_umax.{v, u}, FreeMonoid] end FreeAlgebra + +open Cardinal + +theorem Algebra.rank_adjoin_le {R : Type u} {S : Type v} [CommRing R] [Ring S] [Algebra R S] + (s : Set S) : Module.rank R (adjoin R s) ≤ max #s ℵ₀ := by + rw [adjoin_eq_range_freeAlgebra_lift] + cases subsingleton_or_nontrivial R + · rw [rank_subsingleton]; exact one_le_aleph0.trans (le_max_right _ _) + rw [← lift_le.{max u v}] + refine (lift_rank_range_le (FreeAlgebra.lift R ((↑) : s → S)).toLinearMap).trans ?_ + rw [FreeAlgebra.rank_eq, lift_id'.{v,u}, lift_umax.{v,u}, lift_le, max_comm] + exact mk_list_le_max _ diff --git a/Mathlib/Order/SuccPred/Basic.lean b/Mathlib/Order/SuccPred/Basic.lean index 55b6157d4a5db..ac3781d7a4350 100644 --- a/Mathlib/Order/SuccPred/Basic.lean +++ b/Mathlib/Order/SuccPred/Basic.lean @@ -992,7 +992,7 @@ namespace WithTop section Succ -variable [DecidableEq α] [PartialOrder α] [SuccOrder α] +variable [PartialOrder α] [SuccOrder α] [∀ a : α, Decidable (succ a = a)] instance : SuccOrder (WithTop α) where succ a := @@ -1144,7 +1144,7 @@ end Succ section Pred -variable [DecidableEq α] [PartialOrder α] [PredOrder α] +variable [PartialOrder α] [PredOrder α] [∀ a : α, Decidable (pred a = a)] instance : PredOrder (WithBot α) where pred a := From 159918f85ffa290bd7f3ad8fc1b72a3d41e2d566 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Wed, 20 Nov 2024 13:18:31 +0000 Subject: [PATCH 145/829] feat: the monoidal category structure on homological complexes (#15695) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Let `c : ComplexShape I` with `I` an additive monoid. If `c` is equipped with the data and axioms `c.TensorSigns`, then the category `HomologicalComplex C c` can be eqquiped with a monoidal category structure if `C` is a monoidal category such that `C` has certain coproducts and both left/right tensoring commute with these. In particular, we obtain a monoidal category structure on `ChainComplex C ℕ` when `C` is an additive monoidal category. Co-authored-by: Kim Morrison Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/Algebra/Homology/Monoidal.lean | 325 ++++++++++++++++++ .../CategoryTheory/GradedObject/Monoidal.lean | 19 + 3 files changed, 345 insertions(+) create mode 100644 Mathlib/Algebra/Homology/Monoidal.lean diff --git a/Mathlib.lean b/Mathlib.lean index db0e4160d4515..2c342449a9a42 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -419,6 +419,7 @@ import Mathlib.Algebra.Homology.ImageToKernel import Mathlib.Algebra.Homology.Linear import Mathlib.Algebra.Homology.LocalCohomology import Mathlib.Algebra.Homology.Localization +import Mathlib.Algebra.Homology.Monoidal import Mathlib.Algebra.Homology.Opposite import Mathlib.Algebra.Homology.QuasiIso import Mathlib.Algebra.Homology.Refinements diff --git a/Mathlib/Algebra/Homology/Monoidal.lean b/Mathlib/Algebra/Homology/Monoidal.lean new file mode 100644 index 0000000000000..e0e185b0f0bec --- /dev/null +++ b/Mathlib/Algebra/Homology/Monoidal.lean @@ -0,0 +1,325 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou, Kim Morrison +-/ +import Mathlib.Algebra.Homology.BifunctorAssociator +import Mathlib.Algebra.Homology.Single +import Mathlib.CategoryTheory.GradedObject.Monoidal +import Mathlib.CategoryTheory.Monoidal.Transport + +/-! +# The monoidal category structure on homological complexes + +Let `c : ComplexShape I` with `I` an additive monoid. If `c` is equipped +with the data and axioms `c.TensorSigns`, then the category +`HomologicalComplex C c` can be equipped with a monoidal category +structure if `C` is a monoidal category such that `C` has certain +coproducts and both left/right tensoring commute with these. + +In particular, we obtain a monoidal category structure on +`ChainComplex C ℕ` when `C` is an additive monoidal category. + +-/ + +open CategoryTheory Limits MonoidalCategory Category + +namespace HomologicalComplex + +variable {C : Type*} [Category C] [MonoidalCategory C] [Preadditive C] [HasZeroObject C] + [(curriedTensor C).Additive] [∀ (X₁ : C), ((curriedTensor C).obj X₁).Additive] + {I : Type*} [AddMonoid I] {c : ComplexShape I} [c.TensorSigns] + +/-- If `K₁` and `K₂` are two homological complexes, this is the property that +for all `j`, the coproduct of `K₁ i₁ ⊗ K₂ i₂` for `i₁ + i₂ = j` exists. -/ +abbrev HasTensor (K₁ K₂ : HomologicalComplex C c) := HasMapBifunctor K₁ K₂ (curriedTensor C) c + +section + +variable [DecidableEq I] + +/-- The tensor product of two homological complexes. -/ +noncomputable abbrev tensorObj (K₁ K₂ : HomologicalComplex C c) [HasTensor K₁ K₂] : + HomologicalComplex C c := + mapBifunctor K₁ K₂ (curriedTensor C) c + +/-- The inclusion `K₁.X i₁ ⊗ K₂.X i₂ ⟶ (tensorObj K₁ K₂).X j` of a summand in +the tensor product of the homological complexes. -/ +noncomputable abbrev ιTensorObj (K₁ K₂ : HomologicalComplex C c) [HasTensor K₁ K₂] + (i₁ i₂ j : I) (h : i₁ + i₂ = j) : + K₁.X i₁ ⊗ K₂.X i₂ ⟶ (tensorObj K₁ K₂).X j := + ιMapBifunctor K₁ K₂ (curriedTensor C) c i₁ i₂ j h + +/-- The tensor product of two morphisms of homological complexes. -/ +noncomputable abbrev tensorHom {K₁ K₂ L₁ L₂ : HomologicalComplex C c} + (f : K₁ ⟶ L₁) (g : K₂ ⟶ L₂) [HasTensor K₁ K₂] [HasTensor L₁ L₂] : + tensorObj K₁ K₂ ⟶ tensorObj L₁ L₂ := + mapBifunctorMap f g _ _ + +/-- Given three homological complexes `K₁`, `K₂`, and `K₃`, this asserts that for +all `j`, the functor `- ⊗ K₃.X i₃` commutes with the coproduct of +the `K₁.X i₁ ⊗ K₂.X i₂` such that `i₁ + i₂ = j`. -/ +abbrev HasGoodTensor₁₂ (K₁ K₂ K₃ : HomologicalComplex C c) := + HasGoodTrifunctor₁₂Obj (curriedTensor C) (curriedTensor C) K₁ K₂ K₃ c c + +/-- Given three homological complexes `K₁`, `K₂`, and `K₃`, this asserts that for +all `j`, the functor `K₁.X i₁` commutes with the coproduct of +the `K₂.X i₂ ⊗ K₃.X i₃` such that `i₂ + i₃ = j`. -/ +abbrev HasGoodTensor₂₃ (K₁ K₂ K₃ : HomologicalComplex C c) := + HasGoodTrifunctor₂₃Obj (curriedTensor C) (curriedTensor C) K₁ K₂ K₃ c c c + +/-- The associator isomorphism for the tensor product of homological complexes. -/ +noncomputable abbrev associator (K₁ K₂ K₃ : HomologicalComplex C c) + [HasTensor K₁ K₂] [HasTensor K₂ K₃] + [HasTensor (tensorObj K₁ K₂) K₃] [HasTensor K₁ (tensorObj K₂ K₃)] + [HasGoodTensor₁₂ K₁ K₂ K₃] [HasGoodTensor₂₃ K₁ K₂ K₃] : + tensorObj (tensorObj K₁ K₂) K₃ ≅ tensorObj K₁ (tensorObj K₂ K₃) := + mapBifunctorAssociator (curriedAssociatorNatIso C) K₁ K₂ K₃ c c c + +variable (C c) in +/-- The unit of the tensor product of homological complexes. -/ +noncomputable abbrev tensorUnit : HomologicalComplex C c := (single C c 0).obj (𝟙_ C) + +variable (C c) in +/-- As a graded object, the single complex `(single C c 0).obj (𝟙_ C)` identifies +to the unit `(GradedObject.single₀ I).obj (𝟙_ C)` of the tensor product of graded objects. -/ +noncomputable def tensorUnitIso : + (GradedObject.single₀ I).obj (𝟙_ C) ≅ (tensorUnit C c).X := + GradedObject.isoMk _ _ (fun i ↦ + if hi : i = 0 then + (GradedObject.singleObjApplyIsoOfEq (0 : I) (𝟙_ C) i hi).trans + (singleObjXIsoOfEq c 0 (𝟙_ C) i hi).symm + else + { hom := 0 + inv := 0 + hom_inv_id := (GradedObject.isInitialSingleObjApply 0 (𝟙_ C) i hi).hom_ext _ _ + inv_hom_id := (isZero_single_obj_X c 0 (𝟙_ C) i hi).eq_of_src _ _ }) + +end + +instance (K₁ K₂ : HomologicalComplex C c) [GradedObject.HasTensor K₁.X K₂.X] : + HasTensor K₁ K₂ := by + assumption + +instance (K₁ K₂ K₃ : HomologicalComplex C c) + [GradedObject.HasGoodTensor₁₂Tensor K₁.X K₂.X K₃.X] : + HasGoodTensor₁₂ K₁ K₂ K₃ := + inferInstanceAs (GradedObject.HasGoodTensor₁₂Tensor K₁.X K₂.X K₃.X) + +instance (K₁ K₂ K₃ : HomologicalComplex C c) + [GradedObject.HasGoodTensorTensor₂₃ K₁.X K₂.X K₃.X] : + HasGoodTensor₂₃ K₁ K₂ K₃ := + inferInstanceAs (GradedObject.HasGoodTensorTensor₂₃ K₁.X K₂.X K₃.X) + +section + +variable (K : HomologicalComplex C c) [DecidableEq I] + +section + +variable [∀ X₂, PreservesColimit (Functor.empty.{0} C) ((curriedTensor C).flip.obj X₂)] + +instance : GradedObject.HasTensor (tensorUnit C c).X K.X := + GradedObject.hasTensor_of_iso (tensorUnitIso C c) (Iso.refl _) + +instance : HasTensor (tensorUnit C c) K := + inferInstanceAs (GradedObject.HasTensor (tensorUnit C c).X K.X) + +@[simp] +lemma unit_tensor_d₁ (i₁ i₂ j : I) : + mapBifunctor.d₁ (tensorUnit C c) K (curriedTensor C) c i₁ i₂ j = 0 := by + by_cases h₁ : c.Rel i₁ (c.next i₁) + · by_cases h₂ : ComplexShape.π c c c (c.next i₁, i₂) = j + · rw [mapBifunctor.d₁_eq _ _ _ _ h₁ _ _ h₂, single_obj_d, Functor.map_zero, + zero_app, zero_comp, smul_zero] + · rw [mapBifunctor.d₁_eq_zero' _ _ _ _ h₁ _ _ h₂] + · rw [mapBifunctor.d₁_eq_zero _ _ _ _ _ _ _ h₁] + +end + +section + +variable [∀ X₁, PreservesColimit (Functor.empty.{0} C) ((curriedTensor C).obj X₁)] + +instance : GradedObject.HasTensor K.X (tensorUnit C c).X := + GradedObject.hasTensor_of_iso (Iso.refl _) (tensorUnitIso C c) + +instance : HasTensor K (tensorUnit C c) := + inferInstanceAs (GradedObject.HasTensor K.X (tensorUnit C c).X) + +@[simp] +lemma tensor_unit_d₂ (i₁ i₂ j : I) : + mapBifunctor.d₂ K (tensorUnit C c) (curriedTensor C) c i₁ i₂ j = 0 := by + by_cases h₁ : c.Rel i₂ (c.next i₂) + · by_cases h₂ : ComplexShape.π c c c (i₁, c.next i₂) = j + · rw [mapBifunctor.d₂_eq _ _ _ _ _ h₁ _ h₂, single_obj_d, Functor.map_zero, + zero_comp, smul_zero] + · rw [mapBifunctor.d₂_eq_zero' _ _ _ _ _ h₁ _ h₂] + · rw [mapBifunctor.d₂_eq_zero _ _ _ _ _ _ _ h₁] + +end + +end + +section Unitor + +variable (K : HomologicalComplex C c) [DecidableEq I] + +section LeftUnitor + +variable [∀ X₂, PreservesColimit (Functor.empty.{0} C) ((curriedTensor C).flip.obj X₂)] + +/-- Auxiliary definition for `leftUnitor`. -/ +noncomputable def leftUnitor' : + (tensorObj (tensorUnit C c) K).X ≅ K.X := + GradedObject.Monoidal.tensorIso ((tensorUnitIso C c).symm) (Iso.refl _) ≪≫ + GradedObject.Monoidal.leftUnitor K.X + +lemma leftUnitor'_inv (i : I) : + (leftUnitor' K).inv i = (λ_ (K.X i)).inv ≫ ((singleObjXSelf c 0 (𝟙_ C)).inv ▷ (K.X i)) ≫ + ιTensorObj (tensorUnit C c) K 0 i i (zero_add i) := by + dsimp [leftUnitor'] + rw [GradedObject.Monoidal.leftUnitor_inv_apply, assoc, assoc, Iso.cancel_iso_inv_left, + GradedObject.Monoidal.ι_tensorHom] + dsimp + rw [tensorHom_id, ← comp_whiskerRight_assoc] + congr 2 + rw [← cancel_epi (GradedObject.Monoidal.tensorUnit₀ (I := I)).hom, Iso.hom_inv_id_assoc] + dsimp [tensorUnitIso] + rw [dif_pos rfl] + rfl + +@[reassoc] +lemma leftUnitor'_inv_comm (i j : I) : + (leftUnitor' K).inv i ≫ (tensorObj (tensorUnit C c) K).d i j = + K.d i j ≫ (leftUnitor' K).inv j := by + by_cases hij : c.Rel i j + · simp only [leftUnitor'_inv, assoc, mapBifunctor.d_eq, + Preadditive.comp_add, mapBifunctor.ι_D₁, mapBifunctor.ι_D₂, + unit_tensor_d₁, comp_zero, zero_add] + rw [mapBifunctor.d₂_eq _ _ _ _ _ hij _ (by simp)] + dsimp + simp only [ComplexShape.ε_zero, one_smul, ← whisker_exchange_assoc, + id_whiskerLeft, assoc, Iso.inv_hom_id_assoc] + · simp only [shape _ _ _ hij, comp_zero, zero_comp] + +/-- The left unitor for the tensor product of homological complexes. -/ +noncomputable def leftUnitor : + tensorObj (tensorUnit C c) K ≅ K := + Iso.symm (Hom.isoOfComponents (fun i ↦ (GradedObject.eval i).mapIso (leftUnitor' K).symm) + (fun _ _ _ ↦ leftUnitor'_inv_comm _ _ _)) + +end LeftUnitor + +section RightUnitor + +variable [∀ X₁, PreservesColimit (Functor.empty.{0} C) ((curriedTensor C).obj X₁)] + +/-- Auxiliary definition for `rightUnitor`. -/ +noncomputable def rightUnitor' : + (tensorObj K (tensorUnit C c)).X ≅ K.X := + GradedObject.Monoidal.tensorIso (Iso.refl _) ((tensorUnitIso C c).symm) ≪≫ + GradedObject.Monoidal.rightUnitor K.X + +lemma rightUnitor'_inv (i : I) : + (rightUnitor' K).inv i = (ρ_ (K.X i)).inv ≫ ((K.X i) ◁ (singleObjXSelf c 0 (𝟙_ C)).inv) ≫ + ιTensorObj K (tensorUnit C c) i 0 i (add_zero i) := by + dsimp [rightUnitor'] + rw [GradedObject.Monoidal.rightUnitor_inv_apply, assoc, assoc, Iso.cancel_iso_inv_left, + GradedObject.Monoidal.ι_tensorHom] + dsimp + rw [id_tensorHom, ← MonoidalCategory.whiskerLeft_comp_assoc] + congr 2 + rw [← cancel_epi (GradedObject.Monoidal.tensorUnit₀ (I := I)).hom, Iso.hom_inv_id_assoc] + dsimp [tensorUnitIso] + rw [dif_pos rfl] + rfl + +lemma rightUnitor'_inv_comm (i j : I) : + (rightUnitor' K).inv i ≫ (tensorObj K (tensorUnit C c)).d i j = + K.d i j ≫ (rightUnitor' K).inv j := by + by_cases hij : c.Rel i j + · simp only [rightUnitor'_inv, assoc, mapBifunctor.d_eq, + Preadditive.comp_add, mapBifunctor.ι_D₁, mapBifunctor.ι_D₂, + tensor_unit_d₂, comp_zero, add_zero] + rw [mapBifunctor.d₁_eq _ _ _ _ hij _ _ (by simp)] + dsimp + simp only [one_smul, whisker_exchange_assoc, + MonoidalCategory.whiskerRight_id, assoc, Iso.inv_hom_id_assoc] + · simp only [shape _ _ _ hij, comp_zero, zero_comp] + +/-- The right unitor for the tensor product of homological complexes. -/ +noncomputable def rightUnitor : + tensorObj K (tensorUnit C c) ≅ K := + Iso.symm (Hom.isoOfComponents (fun i ↦ (GradedObject.eval i).mapIso (rightUnitor' K).symm) + (fun _ _ _ ↦ rightUnitor'_inv_comm _ _ _)) + +end RightUnitor + +end Unitor + +variable (C c) [∀ (X₁ X₂ : GradedObject I C), GradedObject.HasTensor X₁ X₂] + [∀ X₁, PreservesColimit (Functor.empty.{0} C) ((curriedTensor C).obj X₁)] + [∀ X₂, PreservesColimit (Functor.empty.{0} C) ((curriedTensor C).flip.obj X₂)] + [∀ (X₁ X₂ X₃ X₄ : GradedObject I C), GradedObject.HasTensor₄ObjExt X₁ X₂ X₃ X₄] + [∀ (X₁ X₂ X₃ : GradedObject I C), GradedObject.HasGoodTensor₁₂Tensor X₁ X₂ X₃] + [∀ (X₁ X₂ X₃ : GradedObject I C), GradedObject.HasGoodTensorTensor₂₃ X₁ X₂ X₃] + [DecidableEq I] + +noncomputable instance monoidalCategoryStruct : + MonoidalCategoryStruct (HomologicalComplex C c) where + tensorObj K₁ K₂ := tensorObj K₁ K₂ + whiskerLeft _ _ _ g := tensorHom (𝟙 _) g + whiskerRight f _ := tensorHom f (𝟙 _) + tensorHom f g := tensorHom f g + tensorUnit := tensorUnit C c + associator K₁ K₂ K₃ := associator K₁ K₂ K₃ + leftUnitor K := leftUnitor K + rightUnitor K := rightUnitor K + +/-- The structure which allows to construct the monoidal category structure +on `HomologicalComplex C c` from the monoidal category structure on +graded objects. -/ +noncomputable def Monoidal.inducingFunctorData : + Monoidal.InducingFunctorData (forget C c) where + μIso _ _ := Iso.refl _ + εIso := tensorUnitIso C c + whiskerLeft_eq K₁ K₂ L₂ g := by + dsimp [forget] + erw [comp_id, id_comp] + rfl + whiskerRight_eq {K₁ L₁} f K₂ := by + dsimp [forget] + erw [comp_id, id_comp] + rfl + tensorHom_eq {K₁ L₁ K₂ L₂} f g := by + dsimp [forget] + erw [comp_id, id_comp] + rfl + associator_eq K₁ K₂ K₃ := by + dsimp [forget] + simp only [tensorHom_id, whiskerRight_tensor, id_whiskerRight, + id_comp, Iso.inv_hom_id, comp_id, assoc] + erw [id_whiskerRight, id_comp, id_comp] + rfl + leftUnitor_eq K := by + dsimp + erw [id_comp] + rfl + rightUnitor_eq K := by + dsimp + rw [assoc] + erw [id_comp] + rfl + +noncomputable instance monoidalCategory : MonoidalCategory (HomologicalComplex C c) := + Monoidal.induced _ (Monoidal.inducingFunctorData C c) + +noncomputable example {D : Type*} [Category D] [Preadditive D] [MonoidalCategory D] + [HasZeroObject D] [HasFiniteCoproducts D] [((curriedTensor D).Additive)] + [∀ (X : D), (((curriedTensor D).obj X).Additive)] + [∀ (X : D), PreservesFiniteCoproducts ((curriedTensor D).obj X)] + [∀ (X : D), PreservesFiniteCoproducts ((curriedTensor D).flip.obj X)] : + MonoidalCategory (ChainComplex D ℕ) := inferInstance + +end HomologicalComplex diff --git a/Mathlib/CategoryTheory/GradedObject/Monoidal.lean b/Mathlib/CategoryTheory/GradedObject/Monoidal.lean index 34da9a269a20b..b10fab11e47d8 100644 --- a/Mathlib/CategoryTheory/GradedObject/Monoidal.lean +++ b/Mathlib/CategoryTheory/GradedObject/Monoidal.lean @@ -38,6 +38,14 @@ the coproduct of the objects `X₁ i ⊗ X₂ j` for `i + j = n` exists. -/ abbrev HasTensor (X₁ X₂ : GradedObject I C) : Prop := HasMap (((mapBifunctor (curriedTensor C) I I).obj X₁).obj X₂) (fun ⟨i, j⟩ => i + j) +lemma hasTensor_of_iso {X₁ X₂ Y₁ Y₂ : GradedObject I C} + (e₁ : X₁ ≅ Y₁) (e₂ : X₂ ≅ Y₂) [HasTensor X₁ X₂] : + HasTensor Y₁ Y₂ := by + let e : ((mapBifunctor (curriedTensor C) I I).obj X₁).obj X₂ ≅ + ((mapBifunctor (curriedTensor C) I I).obj Y₁).obj Y₂ := isoMk _ _ + (fun ⟨i, j⟩ ↦ (eval i).mapIso e₁ ⊗ (eval j).mapIso e₂) + exact hasMap_of_iso e _ + namespace Monoidal /-- The tensor product of two graded objects. -/ @@ -120,6 +128,17 @@ lemma tensor_comp {X₁ X₂ X₃ Y₁ Y₂ Y₃ : GradedObject I C} (f₁ : X apply congr_mapMap simp +/-- The isomorphism `tensorObj X₁ Y₁ ≅ tensorObj X₂ Y₂` induced by isomorphisms of graded +objects `e : X₁ ≅ X₂` and `e' : Y₁ ≅ Y₂`. -/ +@[simps] +noncomputable def tensorIso {X₁ X₂ Y₁ Y₂ : GradedObject I C} (e : X₁ ≅ X₂) (e' : Y₁ ≅ Y₂) + [HasTensor X₁ Y₁] [HasTensor X₂ Y₂] : + tensorObj X₁ Y₁ ≅ tensorObj X₂ Y₂ where + hom := tensorHom e.hom e'.hom + inv := tensorHom e.inv e'.inv + hom_inv_id := by simp only [← tensor_comp, Iso.hom_inv_id, tensor_id] + inv_hom_id := by simp only [← tensor_comp, Iso.inv_hom_id, tensor_id] + lemma tensorHom_def {X₁ X₂ Y₁ Y₂ : GradedObject I C} (f : X₁ ⟶ X₂) (g : Y₁ ⟶ Y₂) [HasTensor X₁ Y₁] [HasTensor X₂ Y₂] [HasTensor X₂ Y₁] : tensorHom f g = whiskerRight f Y₁ ≫ whiskerLeft X₂ g := by From 40b557834ca80e53ba1aed3870919826e1987375 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Wed, 20 Nov 2024 13:18:32 +0000 Subject: [PATCH 146/829] feat: expand API around manifold derivatives (#18850) Notably, prove that basic functions are differentiable in the manifold sense, and compute their derivatives. --- Mathlib/Geometry/Manifold/Algebra/Monoid.lean | 23 +++ Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean | 92 +++++++++ .../Manifold/MFDeriv/SpecificFunctions.lean | 193 +++++++++++++++++- 3 files changed, 303 insertions(+), 5 deletions(-) diff --git a/Mathlib/Geometry/Manifold/Algebra/Monoid.lean b/Mathlib/Geometry/Manifold/Algebra/Monoid.lean index 9202595ec5d94..5439b39e461ed 100644 --- a/Mathlib/Geometry/Manifold/Algebra/Monoid.lean +++ b/Mathlib/Geometry/Manifold/Algebra/Monoid.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Nicolò Cavalleri -/ import Mathlib.Geometry.Manifold.ContMDiffMap +import Mathlib.Geometry.Manifold.MFDeriv.Basic /-! # Smooth monoid @@ -132,6 +133,28 @@ theorem smooth_mul_left {a : G} : Smooth I I fun b : G => a * b := theorem smooth_mul_right {a : G} : Smooth I I fun b : G => b * a := smooth_id.mul smooth_const +theorem contMDiff_mul_left {a : G} : ContMDiff I I n (a * ·) := smooth_mul_left.contMDiff + +theorem contMDiffAt_mul_left {a b : G} : ContMDiffAt I I n (a * ·) b := + contMDiff_mul_left.contMDiffAt + +theorem mdifferentiable_mul_left {a : G} : MDifferentiable I I (a * ·) := + contMDiff_mul_left.mdifferentiable le_rfl + +theorem mdifferentiableAt_mul_left {a b : G} : MDifferentiableAt I I (a * ·) b := + contMDiffAt_mul_left.mdifferentiableAt le_rfl + +theorem contMDiff_mul_right {a : G} : ContMDiff I I n (· * a) := smooth_mul_right.contMDiff + +theorem contMDiffAt_mul_right {a b : G} : ContMDiffAt I I n (· * a) b := + contMDiff_mul_right.contMDiffAt + +theorem mdifferentiable_mul_right {a : G} : MDifferentiable I I (· * a) := + contMDiff_mul_right.mdifferentiable le_rfl + +theorem mdifferentiableAt_mul_right {a b : G} : MDifferentiableAt I I (· * a) b := + contMDiffAt_mul_right.mdifferentiableAt le_rfl + end variable (I) (g h : G) diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean b/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean index eecf755b43be2..eb88f2a6c5c87 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean @@ -257,4 +257,96 @@ theorem mdifferentiableOn_extChartAt_symm : intro y hy exact (mdifferentiableWithinAt_extChartAt_symm hy).mono (extChartAt_target_subset_range x) +/-- The composition of the derivative of `extChartAt` with the derivative of the inverse of +`extChartAt` gives the identity. +Version where the basepoint belongs to `(extChartAt I x).target`. -/ +lemma mfderiv_extChartAt_comp_mfderivWithin_extChartAt_symm {x : M} + {y : E} (hy : y ∈ (extChartAt I x).target) : + (mfderiv I 𝓘(𝕜, E) (extChartAt I x) ((extChartAt I x).symm y)) ∘L + (mfderivWithin 𝓘(𝕜, E) I (extChartAt I x).symm (range I) y) = ContinuousLinearMap.id _ _ := by + have U : UniqueMDiffWithinAt 𝓘(𝕜, E) (range ↑I) y := by + apply I.uniqueMDiffOn + exact extChartAt_target_subset_range x hy + have h'y : (extChartAt I x).symm y ∈ (extChartAt I x).source := (extChartAt I x).map_target hy + have h''y : (extChartAt I x).symm y ∈ (chartAt H x).source := by + rwa [← extChartAt_source (I := I)] + rw [← mfderiv_comp_mfderivWithin]; rotate_left + · apply mdifferentiableAt_extChartAt h''y + · exact mdifferentiableWithinAt_extChartAt_symm hy + · exact U + rw [← mfderivWithin_id U] + apply Filter.EventuallyEq.mfderivWithin_eq U + · filter_upwards [extChartAt_target_mem_nhdsWithin_of_mem hy] with z hz + simp only [Function.comp_def, PartialEquiv.right_inv (extChartAt I x) hz, id_eq] + · simp only [Function.comp_def, PartialEquiv.right_inv (extChartAt I x) hy, id_eq] + +/-- The composition of the derivative of `extChartAt` with the derivative of the inverse of +`extChartAt` gives the identity. +Version where the basepoint belongs to `(extChartAt I x).source`. -/ +lemma mfderiv_extChartAt_comp_mfderivWithin_extChartAt_symm' {x : M} + {y : M} (hy : y ∈ (extChartAt I x).source) : + (mfderiv I 𝓘(𝕜, E) (extChartAt I x) y) ∘L + (mfderivWithin 𝓘(𝕜, E) I (extChartAt I x).symm (range I) (extChartAt I x y)) + = ContinuousLinearMap.id _ _ := by + have : y = (extChartAt I x).symm (extChartAt I x y) := ((extChartAt I x).left_inv hy).symm + convert mfderiv_extChartAt_comp_mfderivWithin_extChartAt_symm ((extChartAt I x).map_source hy) + +/-- The composition of the derivative of the inverse of `extChartAt` with the derivative of +`extChartAt` gives the identity. +Version where the basepoint belongs to `(extChartAt I x).target`. -/ +lemma mfderivWithin_extChartAt_symm_comp_mfderiv_extChartAt + {y : E} (hy : y ∈ (extChartAt I x).target) : + (mfderivWithin 𝓘(𝕜, E) I (extChartAt I x).symm (range I) y) ∘L + (mfderiv I 𝓘(𝕜, E) (extChartAt I x) ((extChartAt I x).symm y)) + = ContinuousLinearMap.id _ _ := by + have h'y : (extChartAt I x).symm y ∈ (extChartAt I x).source := (extChartAt I x).map_target hy + have h''y : (extChartAt I x).symm y ∈ (chartAt H x).source := by + rwa [← extChartAt_source (I := I)] + have U' : UniqueMDiffWithinAt I (extChartAt I x).source ((extChartAt I x).symm y) := + (isOpen_extChartAt_source x).uniqueMDiffWithinAt h'y + have : mfderiv I 𝓘(𝕜, E) (extChartAt I x) ((extChartAt I x).symm y) + = mfderivWithin I 𝓘(𝕜, E) (extChartAt I x) (extChartAt I x).source + ((extChartAt I x).symm y) := by + rw [mfderivWithin_eq_mfderiv U'] + exact mdifferentiableAt_extChartAt h''y + rw [this, ← mfderivWithin_comp_of_eq]; rotate_left + · exact mdifferentiableWithinAt_extChartAt_symm hy + · exact (mdifferentiableAt_extChartAt h''y).mdifferentiableWithinAt + · intro z hz + apply extChartAt_target_subset_range x + exact PartialEquiv.map_source (extChartAt I x) hz + · exact U' + · exact PartialEquiv.right_inv (extChartAt I x) hy + rw [← mfderivWithin_id U'] + apply Filter.EventuallyEq.mfderivWithin_eq U' + · filter_upwards [extChartAt_source_mem_nhdsWithin' h'y] with z hz + simp only [Function.comp_def, PartialEquiv.left_inv (extChartAt I x) hz, id_eq] + · simp only [Function.comp_def, PartialEquiv.right_inv (extChartAt I x) hy, id_eq] + +/-- The composition of the derivative of the inverse of `extChartAt` with the derivative of +`extChartAt` gives the identity. +Version where the basepoint belongs to `(extChartAt I x).source`. -/ +lemma mfderivWithin_extChartAt_symm_comp_mfderiv_extChartAt' + {y : M} (hy : y ∈ (extChartAt I x).source) : + (mfderivWithin 𝓘(𝕜, E) I (extChartAt I x).symm (range I) (extChartAt I x y)) ∘L + (mfderiv I 𝓘(𝕜, E) (extChartAt I x) y) + = ContinuousLinearMap.id _ _ := by + have : y = (extChartAt I x).symm (extChartAt I x y) := ((extChartAt I x).left_inv hy).symm + convert mfderivWithin_extChartAt_symm_comp_mfderiv_extChartAt ((extChartAt I x).map_source hy) + +lemma isInvertible_mfderivWithin_extChartAt_symm {y : E} (hy : y ∈ (extChartAt I x).target) : + (mfderivWithin 𝓘(𝕜, E) I (extChartAt I x).symm (range I) y).IsInvertible := + ContinuousLinearMap.IsInvertible.of_inverse + (mfderivWithin_extChartAt_symm_comp_mfderiv_extChartAt hy) + (mfderiv_extChartAt_comp_mfderivWithin_extChartAt_symm hy) + +lemma isInvertible_mfderiv_extChartAt {y : M} (hy : y ∈ (extChartAt I x).source) : + (mfderiv I 𝓘(𝕜, E) (extChartAt I x) y).IsInvertible := by + have h'y : extChartAt I x y ∈ (extChartAt I x).target := (extChartAt I x).map_source hy + have Z := ContinuousLinearMap.IsInvertible.of_inverse + (mfderiv_extChartAt_comp_mfderivWithin_extChartAt_symm h'y) + (mfderivWithin_extChartAt_symm_comp_mfderiv_extChartAt h'y) + have : (extChartAt I x).symm ((extChartAt I x) y) = y := (extChartAt I x).left_inv hy + rwa [this] at Z + end extChartAt diff --git a/Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean b/Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean index 2bbe312b6c383..63e5592d2c537 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean @@ -26,14 +26,31 @@ section SpecificFunctions /-! ### Differentiability of specific functions -/ -variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] +variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] + -- declare a charted space `M` over the pair `(E, H)`. + {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} {M : Type*} - [TopologicalSpace M] [ChartedSpace H M] {E' : Type*} - [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] + [TopologicalSpace M] [ChartedSpace H M] + -- declare a charted space `M'` over the pair `(E', H')`. + {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] + -- declare a charted space `M''` over the pair `(E'', H'')`. {E'' : Type*} [NormedAddCommGroup E''] [NormedSpace 𝕜 E''] {H'' : Type*} [TopologicalSpace H''] {I'' : ModelWithCorners 𝕜 E'' H''} {M'' : Type*} [TopologicalSpace M''] [ChartedSpace H'' M''] + -- declare a charted space `N` over the pair `(F, G)`. + {F : Type*} + [NormedAddCommGroup F] [NormedSpace 𝕜 F] {G : Type*} [TopologicalSpace G] + {J : ModelWithCorners 𝕜 F G} {N : Type*} [TopologicalSpace N] [ChartedSpace G N] + -- declare a charted space `N'` over the pair `(F', G')`. + {F' : Type*} + [NormedAddCommGroup F'] [NormedSpace 𝕜 F'] {G' : Type*} [TopologicalSpace G'] + {J' : ModelWithCorners 𝕜 F' G'} {N' : Type*} [TopologicalSpace N'] [ChartedSpace G' N'] + -- F₁, F₂, F₃, F₄ are normed spaces + {F₁ : Type*} + [NormedAddCommGroup F₁] [NormedSpace 𝕜 F₁] {F₂ : Type*} [NormedAddCommGroup F₂] + [NormedSpace 𝕜 F₂] {F₃ : Type*} [NormedAddCommGroup F₃] [NormedSpace 𝕜 F₃] {F₄ : Type*} + [NormedAddCommGroup F₄] [NormedSpace 𝕜 F₄] namespace ContinuousLinearMap @@ -301,6 +318,128 @@ theorem mfderivWithin_snd {s : Set (M × M')} {x : M × M'} ContinuousLinearMap.snd 𝕜 (TangentSpace I x.1) (TangentSpace I' x.2) := by rw [MDifferentiable.mfderivWithin mdifferentiableAt_snd hxs]; exact mfderiv_snd +theorem MDifferentiableWithinAt.fst {f : N → M × M'} {s : Set N} {x : N} + (hf : MDifferentiableWithinAt J (I.prod I') f s x) : + MDifferentiableWithinAt J I (fun x => (f x).1) s x := + mdifferentiableAt_fst.comp_mdifferentiableWithinAt x hf + +theorem MDifferentiableAt.fst {f : N → M × M'} {x : N} (hf : MDifferentiableAt J (I.prod I') f x) : + MDifferentiableAt J I (fun x => (f x).1) x := + mdifferentiableAt_fst.comp x hf + +theorem MDifferentiable.fst {f : N → M × M'} (hf : MDifferentiable J (I.prod I') f) : + MDifferentiable J I fun x => (f x).1 := + mdifferentiable_fst.comp hf + +theorem MDifferentiableWithinAt.snd {f : N → M × M'} {s : Set N} {x : N} + (hf : MDifferentiableWithinAt J (I.prod I') f s x) : + MDifferentiableWithinAt J I' (fun x => (f x).2) s x := + mdifferentiableAt_snd.comp_mdifferentiableWithinAt x hf + +theorem MDifferentiableAt.snd {f : N → M × M'} {x : N} (hf : MDifferentiableAt J (I.prod I') f x) : + MDifferentiableAt J I' (fun x => (f x).2) x := + mdifferentiableAt_snd.comp x hf + +theorem MDifferentiable.snd {f : N → M × M'} (hf : MDifferentiable J (I.prod I') f) : + MDifferentiable J I' fun x => (f x).2 := + mdifferentiable_snd.comp hf + +theorem mdifferentiableWithinAt_prod_iff (f : M → M' × N') : + MDifferentiableWithinAt I (I'.prod J') f s x ↔ + MDifferentiableWithinAt I I' (Prod.fst ∘ f) s x + ∧ MDifferentiableWithinAt I J' (Prod.snd ∘ f) s x := + ⟨fun h => ⟨h.fst, h.snd⟩, fun h => h.1.prod_mk h.2⟩ + +theorem mdifferentiableWithinAt_prod_module_iff (f : M → F₁ × F₂) : + MDifferentiableWithinAt I 𝓘(𝕜, F₁ × F₂) f s x ↔ + MDifferentiableWithinAt I 𝓘(𝕜, F₁) (Prod.fst ∘ f) s x ∧ + MDifferentiableWithinAt I 𝓘(𝕜, F₂) (Prod.snd ∘ f) s x := by + rw [modelWithCornersSelf_prod, ← chartedSpaceSelf_prod] + exact mdifferentiableWithinAt_prod_iff f + +theorem mdifferentiableAt_prod_iff (f : M → M' × N') : + MDifferentiableAt I (I'.prod J') f x ↔ + MDifferentiableAt I I' (Prod.fst ∘ f) x ∧ MDifferentiableAt I J' (Prod.snd ∘ f) x := by + simp_rw [← mdifferentiableWithinAt_univ]; exact mdifferentiableWithinAt_prod_iff f + +theorem mdifferentiableAt_prod_module_iff (f : M → F₁ × F₂) : + MDifferentiableAt I 𝓘(𝕜, F₁ × F₂) f x ↔ + MDifferentiableAt I 𝓘(𝕜, F₁) (Prod.fst ∘ f) x + ∧ MDifferentiableAt I 𝓘(𝕜, F₂) (Prod.snd ∘ f) x := by + rw [modelWithCornersSelf_prod, ← chartedSpaceSelf_prod] + exact mdifferentiableAt_prod_iff f + +theorem mdifferentiableOn_prod_iff (f : M → M' × N') : + MDifferentiableOn I (I'.prod J') f s ↔ + MDifferentiableOn I I' (Prod.fst ∘ f) s ∧ MDifferentiableOn I J' (Prod.snd ∘ f) s := + ⟨fun h ↦ ⟨fun x hx ↦ ((mdifferentiableWithinAt_prod_iff f).1 (h x hx)).1, + fun x hx ↦ ((mdifferentiableWithinAt_prod_iff f).1 (h x hx)).2⟩ , + fun h x hx ↦ (mdifferentiableWithinAt_prod_iff f).2 ⟨h.1 x hx, h.2 x hx⟩⟩ + +theorem mdifferentiableOn_prod_module_iff (f : M → F₁ × F₂) : + MDifferentiableOn I 𝓘(𝕜, F₁ × F₂) f s ↔ + MDifferentiableOn I 𝓘(𝕜, F₁) (Prod.fst ∘ f) s + ∧ MDifferentiableOn I 𝓘(𝕜, F₂) (Prod.snd ∘ f) s := by + rw [modelWithCornersSelf_prod, ← chartedSpaceSelf_prod] + exact mdifferentiableOn_prod_iff f + +theorem mdifferentiable_prod_iff (f : M → M' × N') : + MDifferentiable I (I'.prod J') f ↔ + MDifferentiable I I' (Prod.fst ∘ f) ∧ MDifferentiable I J' (Prod.snd ∘ f) := + ⟨fun h => ⟨h.fst, h.snd⟩, fun h => by convert h.1.prod_mk h.2⟩ + +theorem mdifferentiable_prod_module_iff (f : M → F₁ × F₂) : + MDifferentiable I 𝓘(𝕜, F₁ × F₂) f ↔ + MDifferentiable I 𝓘(𝕜, F₁) (Prod.fst ∘ f) ∧ MDifferentiable I 𝓘(𝕜, F₂) (Prod.snd ∘ f) := by + rw [modelWithCornersSelf_prod, ← chartedSpaceSelf_prod] + exact mdifferentiable_prod_iff f + + +section prodMap + +variable {f : M → M'} {g : N → N'} {r : Set N} {y : N} + +/-- The product map of two `C^n` functions within a set at a point is `C^n` +within the product set at the product point. -/ +theorem MDifferentiableWithinAt.prod_map' {p : M × N} (hf : MDifferentiableWithinAt I I' f s p.1) + (hg : MDifferentiableWithinAt J J' g r p.2) : + MDifferentiableWithinAt (I.prod J) (I'.prod J') (Prod.map f g) (s ×ˢ r) p := + (hf.comp p mdifferentiableWithinAt_fst (prod_subset_preimage_fst _ _)).prod_mk <| + hg.comp p mdifferentiableWithinAt_snd (prod_subset_preimage_snd _ _) + +theorem MDifferentiableWithinAt.prod_map (hf : MDifferentiableWithinAt I I' f s x) + (hg : MDifferentiableWithinAt J J' g r y) : + MDifferentiableWithinAt (I.prod J) (I'.prod J') (Prod.map f g) (s ×ˢ r) (x, y) := + MDifferentiableWithinAt.prod_map' hf hg + +theorem MDifferentiableAt.prod_map + (hf : MDifferentiableAt I I' f x) (hg : MDifferentiableAt J J' g y) : + MDifferentiableAt (I.prod J) (I'.prod J') (Prod.map f g) (x, y) := by + rw [← mdifferentiableWithinAt_univ] at * + convert hf.prod_map hg + exact univ_prod_univ.symm + +/-- Variant of `MDifferentiableAt.prod_map` in which the point in the product is given as `p` +instead of a pair `(x, y)`. -/ +theorem MDifferentiableAt.prod_map' {p : M × N} + (hf : MDifferentiableAt I I' f p.1) (hg : MDifferentiableAt J J' g p.2) : + MDifferentiableAt (I.prod J) (I'.prod J') (Prod.map f g) p := by + rcases p with ⟨⟩ + exact hf.prod_map hg + +theorem MDifferentiableOn.prod_map + (hf : MDifferentiableOn I I' f s) (hg : MDifferentiableOn J J' g r) : + MDifferentiableOn (I.prod J) (I'.prod J') (Prod.map f g) (s ×ˢ r) := + (hf.comp mdifferentiableOn_fst (prod_subset_preimage_fst _ _)).prod_mk <| + hg.comp mdifferentiableOn_snd (prod_subset_preimage_snd _ _) + +theorem MDifferentiable.prod_map (hf : MDifferentiable I I' f) (hg : MDifferentiable J J' g) : + MDifferentiable (I.prod J) (I'.prod J') (Prod.map f g) := by + intro p + exact (hf p.1).prod_map' (hg p.2) + +end prodMap + @[simp, mfld_simps] theorem tangentMap_prod_snd {p : TangentBundle (I.prod I') (M × M')} : tangentMap (I.prod I') I' Prod.snd p = ⟨p.proj.2, p.2.2⟩ := by @@ -330,22 +469,30 @@ theorem mfderiv_prod_left {x₀ : M} {y₀ : M'} : refine (mdifferentiableAt_id.mfderiv_prod mdifferentiableAt_const).trans ?_ rw [mfderiv_id, mfderiv_const, ContinuousLinearMap.inl] +theorem tangentMap_prod_left {p : TangentBundle I M} {y₀ : M'} : + tangentMap I (I.prod I') (fun x => (x, y₀)) p = ⟨(p.1, y₀), (p.2, 0)⟩ := by + simp only [tangentMap, mfderiv_prod_left, TotalSpace.mk_inj] + rfl + theorem mfderiv_prod_right {x₀ : M} {y₀ : M'} : mfderiv I' (I.prod I') (fun y => (x₀, y)) y₀ = ContinuousLinearMap.inr 𝕜 (TangentSpace I x₀) (TangentSpace I' y₀) := by refine (mdifferentiableAt_const.mfderiv_prod mdifferentiableAt_id).trans ?_ rw [mfderiv_id, mfderiv_const, ContinuousLinearMap.inr] +theorem tangentMap_prod_right {p : TangentBundle I' M'} {x₀ : M} : + tangentMap I' (I.prod I') (fun y => (x₀, y)) p = ⟨(x₀, p.1), (0, p.2)⟩ := by + simp only [tangentMap, mfderiv_prod_right, TotalSpace.mk_inj] + rfl + /-- The total derivative of a function in two variables is the sum of the partial derivatives. Note that to state this (without casts) we need to be able to see through the definition of `TangentSpace`. -/ theorem mfderiv_prod_eq_add {f : M × M' → M''} {p : M × M'} (hf : MDifferentiableAt (I.prod I') I'' f p) : mfderiv (I.prod I') I'' f p = - show E × E' →L[𝕜] E'' from mfderiv (I.prod I') I'' (fun z : M × M' => f (z.1, p.2)) p + mfderiv (I.prod I') I'' (fun z : M × M' => f (p.1, z.2)) p := by - dsimp only erw [mfderiv_comp_of_eq hf (mdifferentiableAt_fst.prod_mk mdifferentiableAt_const) rfl, mfderiv_comp_of_eq hf (mdifferentiableAt_const.prod_mk mdifferentiableAt_snd) rfl, ← ContinuousLinearMap.comp_add, @@ -356,6 +503,42 @@ theorem mfderiv_prod_eq_add {f : M × M' → M''} {p : M × M'} convert ContinuousLinearMap.comp_id <| mfderiv (.prod I I') I'' f (p.1, p.2) exact ContinuousLinearMap.coprod_inl_inr +/-- The total derivative of a function in two variables is the sum of the partial derivatives. + Note that to state this (without casts) we need to be able to see through the definition of + `TangentSpace`. Version in terms of the one-variable derivatives. -/ +theorem mfderiv_prod_eq_add_comp {f : M × M' → M''} {p : M × M'} + (hf : MDifferentiableAt (I.prod I') I'' f p) : + mfderiv (I.prod I') I'' f p = + (mfderiv I I'' (fun z : M => f (z, p.2)) p.1) ∘L (id (ContinuousLinearMap.fst 𝕜 E E') : + (TangentSpace (I.prod I') p) →L[𝕜] (TangentSpace I p.1)) + + (mfderiv I' I'' (fun z : M' => f (p.1, z)) p.2) ∘L (id (ContinuousLinearMap.snd 𝕜 E E') : + (TangentSpace (I.prod I') p) →L[𝕜] (TangentSpace I' p.2)) := by + rw [mfderiv_prod_eq_add hf] + congr + · have : (fun z : M × M' => f (z.1, p.2)) = (fun z : M => f (z, p.2)) ∘ Prod.fst := rfl + rw [this, mfderiv_comp (I' := I)] + · simp only [mfderiv_fst, id_eq] + rfl + · exact hf.comp _ (mdifferentiableAt_id.prod_mk mdifferentiableAt_const) + · exact mdifferentiableAt_fst + · have : (fun z : M × M' => f (p.1, z.2)) = (fun z : M' => f (p.1, z)) ∘ Prod.snd := rfl + rw [this, mfderiv_comp (I' := I')] + · simp only [mfderiv_snd, id_eq] + rfl + · exact hf.comp _ (mdifferentiableAt_const.prod_mk mdifferentiableAt_id) + · exact mdifferentiableAt_snd + +/-- The total derivative of a function in two variables is the sum of the partial derivatives. + Note that to state this (without casts) we need to be able to see through the definition of + `TangentSpace`. Version in terms of the one-variable derivatives. -/ +theorem mfderiv_prod_eq_add_apply {f : M × M' → M''} {p : M × M'} {v : TangentSpace (I.prod I') p} + (hf : MDifferentiableAt (I.prod I') I'' f p) : + mfderiv (I.prod I') I'' f p v = + mfderiv I I'' (fun z : M => f (z, p.2)) p.1 v.1 + + mfderiv I' I'' (fun z : M' => f (p.1, z)) p.2 v.2 := by + rw [mfderiv_prod_eq_add_comp hf] + rfl + end Prod section Arithmetic From fc5f4490ca3100759e2ad673a8e67ea5ad23b3f5 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 20 Nov 2024 13:18:33 +0000 Subject: [PATCH 147/829] ci: print the existing update PR, if found (#19218) This reveals that the problem was #18946 being open. --- .github/workflows/update_dependencies.yml | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/.github/workflows/update_dependencies.yml b/.github/workflows/update_dependencies.yml index 48111dc1fac1d..a8ea709c6de9e 100644 --- a/.github/workflows/update_dependencies.yml +++ b/.github/workflows/update_dependencies.yml @@ -40,6 +40,13 @@ jobs: # Only return if PR is still open filterOutClosed: true + - name: Print PR, if found + run: echo "Found PR ${prNumber} at ${prUrl}" + if: steps.PR.outputs.pr_found == 'true' + env: + prNumber: ${{ steps.PR.outputs.number }} + prUrl: ${{ steps.PR.outputs.pr_url }} + - name: Configure Git User if: ${{ !contains(steps.PR.outputs.pr_labels, 'ready-to-merge') }} run: | From 025632cce0c65284113a5c478a7c6303fdb92a49 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Wed, 20 Nov 2024 13:28:38 +0000 Subject: [PATCH 148/829] =?UTF-8?q?feat:=20change=20definition=20of=20`Has?= =?UTF-8?q?FTaylorSeries`=20to=20take=20a=20parameter=20in=20`WithTop=20?= =?UTF-8?q?=E2=84=95=E2=88=9E`=20instead=20of=20`=E2=84=95=E2=88=9E`=20(#1?= =?UTF-8?q?8723)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit For now this is useless (and even counter-productive, as it means we have to add some casts in some places), but it will make it possible to integrate analytic functions in the smooth hierarchy in #17152. --- Mathlib/Analysis/Calculus/ContDiff/Basic.lean | 49 +++---- Mathlib/Analysis/Calculus/ContDiff/Defs.lean | 101 +++++++------- .../Calculus/ContDiff/FTaylorSeries.lean | 128 ++++++++++++------ .../Calculus/ContDiff/FaaDiBruno.lean | 14 +- .../Calculus/ContDiff/FiniteDimension.lean | 6 +- .../Analysis/Calculus/ContDiff/RCLike.lean | 3 +- .../Analysis/Calculus/FDeriv/Analytic.lean | 2 +- Mathlib/Analysis/Convolution.lean | 3 +- .../Fourier/FourierTransformDeriv.lean | 22 ++- .../Analysis/SpecialFunctions/Pow/Deriv.lean | 2 +- Mathlib/Data/ENat/Basic.lean | 25 +++- 11 files changed, 216 insertions(+), 139 deletions(-) diff --git a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean index 05d77c24f110e..52ac217ea8cf2 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean @@ -187,7 +187,7 @@ theorem IsBoundedBilinearMap.contDiff (hb : IsBoundedBilinearMap 𝕜 b) : ContD /-- If `f` admits a Taylor series `p` in a set `s`, and `g` is linear, then `g ∘ f` admits a Taylor series whose `k`-th term is given by `g ∘ (p k)`. -/ -theorem HasFTaylorSeriesUpToOn.continuousLinearMap_comp (g : F →L[𝕜] G) +theorem HasFTaylorSeriesUpToOn.continuousLinearMap_comp {n : WithTop ℕ∞} (g : F →L[𝕜] G) (hf : HasFTaylorSeriesUpToOn n f p s) : HasFTaylorSeriesUpToOn n (g ∘ f) (fun x k => g.compContinuousMultilinearMap (p x k)) s where zero_eq x hx := congr_arg g (hf.zero_eq x hx) @@ -221,16 +221,16 @@ theorem ContDiff.continuousLinearMap_comp {f : E → F} (g : F →L[𝕜] G) (hf /-- The iterated derivative within a set of the composition with a linear map on the left is obtained by applying the linear map to the iterated derivative. -/ theorem ContinuousLinearMap.iteratedFDerivWithin_comp_left {f : E → F} (g : F →L[𝕜] G) - (hf : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {i : ℕ} (hi : (i : ℕ∞) ≤ n) : + (hf : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {i : ℕ} (hi : i ≤ n) : iteratedFDerivWithin 𝕜 i (g ∘ f) s x = g.compContinuousMultilinearMap (iteratedFDerivWithin 𝕜 i f s x) := (((hf.ftaylorSeriesWithin hs).continuousLinearMap_comp g).eq_iteratedFDerivWithin_of_uniqueDiffOn - hi hs hx).symm + (mod_cast hi) hs hx).symm /-- The iterated derivative of the composition with a linear map on the left is obtained by applying the linear map to the iterated derivative. -/ theorem ContinuousLinearMap.iteratedFDeriv_comp_left {f : E → F} (g : F →L[𝕜] G) - (hf : ContDiff 𝕜 n f) (x : E) {i : ℕ} (hi : (i : ℕ∞) ≤ n) : + (hf : ContDiff 𝕜 n f) (x : E) {i : ℕ} (hi : i ≤ n) : iteratedFDeriv 𝕜 i (g ∘ f) x = g.compContinuousMultilinearMap (iteratedFDeriv 𝕜 i f x) := by simp only [← iteratedFDerivWithin_univ] exact g.iteratedFDerivWithin_comp_left hf.contDiffOn uniqueDiffOn_univ (mem_univ x) hi @@ -262,7 +262,7 @@ theorem ContinuousLinearEquiv.iteratedFDerivWithin_comp_left (g : F ≃L[𝕜] G /-- Composition with a linear isometry on the left preserves the norm of the iterated derivative within a set. -/ theorem LinearIsometry.norm_iteratedFDerivWithin_comp_left {f : E → F} (g : F →ₗᵢ[𝕜] G) - (hf : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {i : ℕ} (hi : (i : ℕ∞) ≤ n) : + (hf : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {i : ℕ} (hi : i ≤ n) : ‖iteratedFDerivWithin 𝕜 i (g ∘ f) s x‖ = ‖iteratedFDerivWithin 𝕜 i f s x‖ := by have : iteratedFDerivWithin 𝕜 i (g ∘ f) s x = @@ -274,7 +274,7 @@ theorem LinearIsometry.norm_iteratedFDerivWithin_comp_left {f : E → F} (g : F /-- Composition with a linear isometry on the left preserves the norm of the iterated derivative. -/ theorem LinearIsometry.norm_iteratedFDeriv_comp_left {f : E → F} (g : F →ₗᵢ[𝕜] G) - (hf : ContDiff 𝕜 n f) (x : E) {i : ℕ} (hi : (i : ℕ∞) ≤ n) : + (hf : ContDiff 𝕜 n f) (x : E) {i : ℕ} (hi : i ≤ n) : ‖iteratedFDeriv 𝕜 i (g ∘ f) x‖ = ‖iteratedFDeriv 𝕜 i f x‖ := by simp only [← iteratedFDerivWithin_univ] exact g.norm_iteratedFDerivWithin_comp_left hf.contDiffOn uniqueDiffOn_univ (mem_univ x) hi @@ -326,8 +326,8 @@ theorem ContinuousLinearEquiv.comp_contDiff_iff (e : F ≃L[𝕜] G) : /-- If `f` admits a Taylor series `p` in a set `s`, and `g` is linear, then `f ∘ g` admits a Taylor series in `g ⁻¹' s`, whose `k`-th term is given by `p k (g v₁, ..., g vₖ)` . -/ -theorem HasFTaylorSeriesUpToOn.compContinuousLinearMap (hf : HasFTaylorSeriesUpToOn n f p s) - (g : G →L[𝕜] E) : +theorem HasFTaylorSeriesUpToOn.compContinuousLinearMap {n : WithTop ℕ∞} + (hf : HasFTaylorSeriesUpToOn n f p s) (g : G →L[𝕜] E) : HasFTaylorSeriesUpToOn n (f ∘ g) (fun x k => (p (g x) k).compContinuousLinearMap fun _ => g) (g ⁻¹' s) := by let A : ∀ m : ℕ, (E[×m]→L[𝕜] F) → G[×m]→L[𝕜] F := fun m h => h.compContinuousLinearMap fun _ => g @@ -372,11 +372,11 @@ theorem ContDiff.comp_continuousLinearMap {f : E → F} {g : G →L[𝕜] E} (hf obtained by composing the iterated derivative with the linear map. -/ theorem ContinuousLinearMap.iteratedFDerivWithin_comp_right {f : E → F} (g : G →L[𝕜] E) (hf : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s) (h's : UniqueDiffOn 𝕜 (g ⁻¹' s)) {x : G} - (hx : g x ∈ s) {i : ℕ} (hi : (i : ℕ∞) ≤ n) : + (hx : g x ∈ s) {i : ℕ} (hi : i ≤ n) : iteratedFDerivWithin 𝕜 i (f ∘ g) (g ⁻¹' s) x = (iteratedFDerivWithin 𝕜 i f s (g x)).compContinuousLinearMap fun _ => g := (((hf.ftaylorSeriesWithin hs).compContinuousLinearMap g).eq_iteratedFDerivWithin_of_uniqueDiffOn - hi h's hx).symm + (mod_cast hi) h's hx).symm /-- The iterated derivative within a set of the composition with a linear equiv on the right is obtained by composing the iterated derivative with the linear equiv. -/ @@ -406,7 +406,7 @@ theorem ContinuousLinearEquiv.iteratedFDerivWithin_comp_right (g : G ≃L[𝕜] /-- The iterated derivative of the composition with a linear map on the right is obtained by composing the iterated derivative with the linear map. -/ theorem ContinuousLinearMap.iteratedFDeriv_comp_right (g : G →L[𝕜] E) {f : E → F} - (hf : ContDiff 𝕜 n f) (x : G) {i : ℕ} (hi : (i : ℕ∞) ≤ n) : + (hf : ContDiff 𝕜 n f) (x : G) {i : ℕ} (hi : i ≤ n) : iteratedFDeriv 𝕜 i (f ∘ g) x = (iteratedFDeriv 𝕜 i f (g x)).compContinuousLinearMap fun _ => g := by simp only [← iteratedFDerivWithin_univ] @@ -463,7 +463,8 @@ theorem ContinuousLinearEquiv.contDiff_comp_iff (e : G ≃L[𝕜] E) : /-- If two functions `f` and `g` admit Taylor series `p` and `q` in a set `s`, then the cartesian product of `f` and `g` admits the cartesian product of `p` and `q` as a Taylor series. -/ -theorem HasFTaylorSeriesUpToOn.prod (hf : HasFTaylorSeriesUpToOn n f p s) {g : E → G} +theorem HasFTaylorSeriesUpToOn.prod {n : WithTop ℕ∞} + (hf : HasFTaylorSeriesUpToOn n f p s) {g : E → G} {q : E → FormalMultilinearSeries 𝕜 E G} (hg : HasFTaylorSeriesUpToOn n g q s) : HasFTaylorSeriesUpToOn n (fun y => (f y, g y)) (fun y k => (p y k).prod (q y k)) s := by set L := fun m => ContinuousMultilinearMap.prodL 𝕜 (fun _ : Fin m => E) F G @@ -1131,7 +1132,7 @@ variable {ι ι' : Type*} [Fintype ι] [Fintype ι'] {F' : ι → Type*} [∀ i, [∀ i, NormedSpace 𝕜 (F' i)] {φ : ∀ i, E → F' i} {p' : ∀ i, E → FormalMultilinearSeries 𝕜 E (F' i)} {Φ : E → ∀ i, F' i} {P' : E → FormalMultilinearSeries 𝕜 E (∀ i, F' i)} -theorem hasFTaylorSeriesUpToOn_pi : +theorem hasFTaylorSeriesUpToOn_pi {n : WithTop ℕ∞} : HasFTaylorSeriesUpToOn n (fun x i => φ i x) (fun x m => ContinuousMultilinearMap.pi fun i => p' i x m) s ↔ ∀ i, HasFTaylorSeriesUpToOn n (φ i) (p' i) s := by @@ -1150,7 +1151,7 @@ theorem hasFTaylorSeriesUpToOn_pi : exact (L m).continuous.comp_continuousOn <| continuousOn_pi.2 fun i => (h i).cont m hm @[simp] -theorem hasFTaylorSeriesUpToOn_pi' : +theorem hasFTaylorSeriesUpToOn_pi' {n : WithTop ℕ∞} : HasFTaylorSeriesUpToOn n Φ P' s ↔ ∀ i, HasFTaylorSeriesUpToOn n (fun x => Φ x i) (fun x m => (@ContinuousLinearMap.proj 𝕜 _ ι F' _ _ _ i).compContinuousMultilinearMap @@ -1204,7 +1205,7 @@ end Pi section Add -theorem HasFTaylorSeriesUpToOn.add {q g} (hf : HasFTaylorSeriesUpToOn n f p s) +theorem HasFTaylorSeriesUpToOn.add {n : WithTop ℕ∞} {q g} (hf : HasFTaylorSeriesUpToOn n f p s) (hg : HasFTaylorSeriesUpToOn n g q s) : HasFTaylorSeriesUpToOn n (f + g) (p + q) s := by convert HasFTaylorSeriesUpToOn.continuousLinearMap_comp (ContinuousLinearMap.fst 𝕜 F F + .snd 𝕜 F F) (hf.prod hg) @@ -1641,12 +1642,13 @@ theorem contDiffAt_ring_inverse [CompleteSpace R] (x : Rˣ) : refine ⟨{ y : R | IsUnit y }, ?_, ?_⟩ · simpa [nhdsWithin_univ] using x.nhds · use ftaylorSeriesWithin 𝕜 inverse univ - rw [le_antisymm hm bot_le, hasFTaylorSeriesUpToOn_zero_iff] + have : (m : WithTop ℕ∞) = 0 := by exact_mod_cast le_antisymm hm bot_le + rw [this, hasFTaylorSeriesUpToOn_zero_iff] constructor · rintro _ ⟨x', rfl⟩ exact (inverse_continuousAt x').continuousWithinAt · simp [ftaylorSeriesWithin] - · rw [contDiffAt_succ_iff_hasFDerivAt] + · rw [show (n.succ : ℕ∞) = n + 1 from rfl, contDiffAt_succ_iff_hasFDerivAt] refine ⟨fun x : R => -mulLeftRight 𝕜 R (inverse x) (inverse x), ?_, ?_⟩ · refine ⟨{ y : R | IsUnit y }, x.nhds, ?_⟩ rintro _ ⟨y, rfl⟩ @@ -1752,7 +1754,7 @@ theorem PartialHomeomorph.contDiffAt_symm [CompleteSpace E] (f : PartialHomeomor · rw [contDiffAt_zero] exact ⟨f.target, IsOpen.mem_nhds f.open_target ha, f.continuousOn_invFun⟩ · obtain ⟨f', ⟨u, hu, hff'⟩, hf'⟩ := contDiffAt_succ_iff_hasFDerivAt.mp hf - rw [contDiffAt_succ_iff_hasFDerivAt] + rw [show (n.succ : ℕ∞) = n + 1 from rfl, contDiffAt_succ_iff_hasFDerivAt] -- For showing `n.succ` times continuous differentiability (the main inductive step), it -- suffices to produce the derivative and show that it is `n` times continuously differentiable have eq_f₀' : f' (f.symm a) = f₀' := (hff' (f.symm a) (mem_of_mem_nhds hu)).unique hf₀' @@ -1871,7 +1873,7 @@ open ContinuousLinearMap (smulRight) /-- A function is `C^(n + 1)` on a domain with unique derivatives if and only if it is differentiable there, and its derivative (formulated with `derivWithin`) is `C^n`. -/ theorem contDiffOn_succ_iff_derivWithin {n : ℕ} (hs : UniqueDiffOn 𝕜 s₂) : - ContDiffOn 𝕜 (n + 1 : ℕ) f₂ s₂ ↔ + ContDiffOn 𝕜 (n + 1) f₂ s₂ ↔ DifferentiableOn 𝕜 f₂ s₂ ∧ ContDiffOn 𝕜 n (derivWithin f₂ s₂) s₂ := by rw [contDiffOn_succ_iff_fderivWithin hs, and_congr_right_iff] intro _ @@ -1893,7 +1895,7 @@ theorem contDiffOn_succ_iff_derivWithin {n : ℕ} (hs : UniqueDiffOn 𝕜 s₂) /-- A function is `C^(n + 1)` on an open domain if and only if it is differentiable there, and its derivative (formulated with `deriv`) is `C^n`. -/ theorem contDiffOn_succ_iff_deriv_of_isOpen {n : ℕ} (hs : IsOpen s₂) : - ContDiffOn 𝕜 (n + 1 : ℕ) f₂ s₂ ↔ DifferentiableOn 𝕜 f₂ s₂ ∧ ContDiffOn 𝕜 n (deriv f₂) s₂ := by + ContDiffOn 𝕜 (n + 1) f₂ s₂ ↔ DifferentiableOn 𝕜 f₂ s₂ ∧ ContDiffOn 𝕜 n (deriv f₂) s₂ := by rw [contDiffOn_succ_iff_derivWithin hs.uniqueDiffOn] exact Iff.rfl.and (contDiffOn_congr fun _ => derivWithin_of_isOpen hs) @@ -1944,7 +1946,7 @@ theorem ContDiffOn.continuousOn_deriv_of_isOpen (h : ContDiffOn 𝕜 n f₂ s₂ /-- A function is `C^(n + 1)` if and only if it is differentiable, and its derivative (formulated in terms of `deriv`) is `C^n`. -/ theorem contDiff_succ_iff_deriv {n : ℕ} : - ContDiff 𝕜 (n + 1 : ℕ) f₂ ↔ Differentiable 𝕜 f₂ ∧ ContDiff 𝕜 n (deriv f₂) := by + ContDiff 𝕜 (n + 1) f₂ ↔ Differentiable 𝕜 f₂ ∧ ContDiff 𝕜 n (deriv f₂) := by simp only [← contDiffOn_univ, contDiffOn_succ_iff_deriv_of_isOpen, isOpen_univ, differentiableOn_univ] @@ -1992,7 +1994,8 @@ variable [NormedSpace 𝕜' E] [IsScalarTower 𝕜 𝕜' E] variable [NormedSpace 𝕜' F] [IsScalarTower 𝕜 𝕜' F] variable {p' : E → FormalMultilinearSeries 𝕜' E F} -theorem HasFTaylorSeriesUpToOn.restrictScalars (h : HasFTaylorSeriesUpToOn n f p' s) : +theorem HasFTaylorSeriesUpToOn.restrictScalars {n : WithTop ℕ∞} + (h : HasFTaylorSeriesUpToOn n f p' s) : HasFTaylorSeriesUpToOn n f (fun x => (p' x).restrictScalars 𝕜) s where zero_eq x hx := h.zero_eq x hx fderivWithin m hm x hx := by @@ -2017,4 +2020,4 @@ theorem ContDiff.restrict_scalars (h : ContDiff 𝕜' n f) : ContDiff 𝕜 n f : end RestrictScalars -set_option linter.style.longFile 2100 +set_option linter.style.longFile 2200 diff --git a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean index c7f8b24aa799b..968b9b7d7d1c3 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean @@ -124,7 +124,7 @@ For instance, a real function which is `C^m` on `(-1/m, 1/m)` for each natural ` better, is `C^∞` at `0` within `univ`. -/ def ContDiffWithinAt (n : ℕ∞) (f : E → F) (s : Set E) (x : E) : Prop := - ∀ m : ℕ, (m : ℕ∞) ≤ n → ∃ u ∈ 𝓝[insert x s] x, + ∀ m : ℕ, m ≤ n → ∃ u ∈ 𝓝[insert x s] x, ∃ p : E → FormalMultilinearSeries 𝕜 E F, HasFTaylorSeriesUpToOn m f p u variable {𝕜} @@ -132,7 +132,7 @@ variable {𝕜} theorem contDiffWithinAt_nat {n : ℕ} : ContDiffWithinAt 𝕜 n f s x ↔ ∃ u ∈ 𝓝[insert x s] x, ∃ p : E → FormalMultilinearSeries 𝕜 E F, HasFTaylorSeriesUpToOn n f p u := - ⟨fun H => H n le_rfl, fun ⟨u, hu, p, hp⟩ _m hm => ⟨u, hu, p, hp.of_le hm⟩⟩ + ⟨fun H => H n le_rfl, fun ⟨u, hu, p, hp⟩ _m hm => ⟨u, hu, p, hp.of_le (mod_cast hm)⟩⟩ theorem ContDiffWithinAt.of_le (h : ContDiffWithinAt 𝕜 n f s x) (hmn : m ≤ n) : ContDiffWithinAt 𝕜 m f s x := fun k hk => h k (le_trans hk hmn) @@ -290,30 +290,32 @@ theorem ContDiffWithinAt.differentiableWithinAt (h : ContDiffWithinAt 𝕜 n f s /-- A function is `C^(n + 1)` on a domain iff locally, it has a derivative which is `C^n`. -/ theorem contDiffWithinAt_succ_iff_hasFDerivWithinAt {n : ℕ} : - ContDiffWithinAt 𝕜 (n + 1 : ℕ) f s x ↔ ∃ u ∈ 𝓝[insert x s] x, ∃ f' : E → E →L[𝕜] F, + ContDiffWithinAt 𝕜 (n + 1) f s x ↔ ∃ u ∈ 𝓝[insert x s] x, ∃ f' : E → E →L[𝕜] F, (∀ x ∈ u, HasFDerivWithinAt f (f' x) u x) ∧ ContDiffWithinAt 𝕜 n f' u x := by constructor · intro h rcases h n.succ le_rfl with ⟨u, hu, p, Hp⟩ refine ⟨u, hu, fun y => (continuousMultilinearCurryFin1 𝕜 E F) (p y 1), fun y hy => - Hp.hasFDerivWithinAt (WithTop.coe_le_coe.2 (Nat.le_add_left 1 n)) hy, ?_⟩ + Hp.hasFDerivWithinAt (mod_cast (Nat.le_add_left 1 n)) hy, ?_⟩ intro m hm refine ⟨u, ?_, fun y : E => (p y).shift, ?_⟩ · -- Porting note: without the explicit argument Lean is not sure of the type. convert @self_mem_nhdsWithin _ _ x u have : x ∈ insert x s := by simp exact insert_eq_of_mem (mem_of_mem_nhdsWithin this hu) - · rw [hasFTaylorSeriesUpToOn_succ_iff_right] at Hp - exact Hp.2.2.of_le hm + · rw [show ((n.succ : ℕ) : WithTop ℕ∞) = n + 1 from rfl, + hasFTaylorSeriesUpToOn_succ_iff_right] at Hp + exact Hp.2.2.of_le (mod_cast hm) · rintro ⟨u, hu, f', f'_eq_deriv, Hf'⟩ - rw [contDiffWithinAt_nat] + rw [show (n : ℕ∞) + 1 = (n + 1 : ℕ) from rfl, contDiffWithinAt_nat] rcases Hf' n le_rfl with ⟨v, hv, p', Hp'⟩ refine ⟨v ∩ u, ?_, fun x => (p' x).unshift (f x), ?_⟩ · apply Filter.inter_mem _ hu apply nhdsWithin_le_of_mem hu exact nhdsWithin_mono _ (subset_insert x u) hv - · rw [hasFTaylorSeriesUpToOn_succ_iff_right] + · rw [show ((n.succ : ℕ) : WithTop ℕ∞) = n + 1 from rfl, + hasFTaylorSeriesUpToOn_succ_iff_right] refine ⟨fun y _ => rfl, fun y hy => ?_, ?_⟩ · change HasFDerivWithinAt (fun z => (continuousMultilinearCurryFin0 𝕜 E F).symm (f z)) @@ -340,7 +342,7 @@ theorem contDiffWithinAt_succ_iff_hasFDerivWithinAt {n : ℕ} : /-- A version of `contDiffWithinAt_succ_iff_hasFDerivWithinAt` where all derivatives are taken within the same set. -/ theorem contDiffWithinAt_succ_iff_hasFDerivWithinAt' {n : ℕ} : - ContDiffWithinAt 𝕜 (n + 1 : ℕ) f s x ↔ + ContDiffWithinAt 𝕜 (n + 1) f s x ↔ ∃ u ∈ 𝓝[insert x s] x, u ⊆ insert x s ∧ ∃ f' : E → E →L[𝕜] F, (∀ x ∈ u, HasFDerivWithinAt f (f' x) s x) ∧ ContDiffWithinAt 𝕜 n f' s x := by refine ⟨fun hf => ?_, ?_⟩ @@ -378,13 +380,13 @@ theorem HasFTaylorSeriesUpToOn.contDiffOn {f' : E → FormalMultilinearSeries intro x hx m hm use s simp only [Set.insert_eq_of_mem hx, self_mem_nhdsWithin, true_and] - exact ⟨f', hf.of_le hm⟩ + exact ⟨f', hf.of_le (mod_cast hm)⟩ theorem ContDiffOn.contDiffWithinAt (h : ContDiffOn 𝕜 n f s) (hx : x ∈ s) : ContDiffWithinAt 𝕜 n f s x := h x hx -theorem ContDiffWithinAt.contDiffOn' {m : ℕ} (hm : (m : ℕ∞) ≤ n) +theorem ContDiffWithinAt.contDiffOn' {m : ℕ} (hm : m ≤ n) (h : ContDiffWithinAt 𝕜 n f s x) : ∃ u, IsOpen u ∧ x ∈ u ∧ ContDiffOn 𝕜 m f (insert x s ∩ u) := by rcases h m hm with ⟨t, ht, p, hp⟩ @@ -392,7 +394,7 @@ theorem ContDiffWithinAt.contDiffOn' {m : ℕ} (hm : (m : ℕ∞) ≤ n) rw [inter_comm] at hut exact ⟨u, huo, hxu, (hp.mono hut).contDiffOn⟩ -theorem ContDiffWithinAt.contDiffOn {m : ℕ} (hm : (m : ℕ∞) ≤ n) (h : ContDiffWithinAt 𝕜 n f s x) : +theorem ContDiffWithinAt.contDiffOn {m : ℕ} (hm : m ≤ n) (h : ContDiffWithinAt 𝕜 n f s x) : ∃ u ∈ 𝓝[insert x s] x, u ⊆ insert x s ∧ ContDiffOn 𝕜 m f u := let ⟨_u, uo, xu, h⟩ := h.contDiffOn' hm ⟨_, inter_mem_nhdsWithin _ (uo.mem_nhds xu), inter_subset_left, h⟩ @@ -466,7 +468,7 @@ theorem contDiffOn_of_locally_contDiffOn /-- A function is `C^(n + 1)` on a domain iff locally, it has a derivative which is `C^n`. -/ theorem contDiffOn_succ_iff_hasFDerivWithinAt {n : ℕ} : - ContDiffOn 𝕜 (n + 1 : ℕ) f s ↔ + ContDiffOn 𝕜 (n + 1) f s ↔ ∀ x ∈ s, ∃ u ∈ 𝓝[insert x s] x, ∃ f' : E → E →L[𝕜] F, (∀ x ∈ u, HasFDerivWithinAt f (f' x) u x) ∧ ContDiffOn 𝕜 n f' u := by constructor @@ -474,10 +476,11 @@ theorem contDiffOn_succ_iff_hasFDerivWithinAt {n : ℕ} : rcases (h x hx) n.succ le_rfl with ⟨u, hu, p, Hp⟩ refine ⟨u, hu, fun y => (continuousMultilinearCurryFin1 𝕜 E F) (p y 1), fun y hy => - Hp.hasFDerivWithinAt (WithTop.coe_le_coe.2 (Nat.le_add_left 1 n)) hy, ?_⟩ - rw [hasFTaylorSeriesUpToOn_succ_iff_right] at Hp + Hp.hasFDerivWithinAt (mod_cast (Nat.le_add_left 1 n)) hy, ?_⟩ + rw [show (n.succ : WithTop ℕ∞) = (n : ℕ) + 1 from rfl, + hasFTaylorSeriesUpToOn_succ_iff_right] at Hp intro z hz m hm - refine ⟨u, ?_, fun x : E => (p x).shift, Hp.2.2.of_le hm⟩ + refine ⟨u, ?_, fun x : E => (p x).shift, Hp.2.2.of_le (mod_cast hm)⟩ -- Porting note: without the explicit arguments `convert` can not determine the type. convert @self_mem_nhdsWithin _ _ z u exact insert_eq_of_mem hz @@ -490,7 +493,7 @@ theorem contDiffOn_succ_iff_hasFDerivWithinAt {n : ℕ} : @[simp] theorem contDiffOn_zero : ContDiffOn 𝕜 0 f s ↔ ContinuousOn f s := by refine ⟨fun H => H.continuousOn, fun H => fun x hx m hm ↦ ?_⟩ - have : (m : ℕ∞) = 0 := le_antisymm hm bot_le + have : (m : WithTop ℕ∞) = 0 := le_antisymm (mod_cast hm) bot_le rw [this] refine ⟨insert x s, self_mem_nhdsWithin, ftaylorSeriesWithin 𝕜 f s, ?_⟩ rw [hasFTaylorSeriesUpToOn_zero_iff] @@ -519,7 +522,8 @@ protected theorem ContDiffOn.ftaylorSeriesWithin (h : ContDiffOn 𝕜 n f s) (hs simp only [ftaylorSeriesWithin, ContinuousMultilinearMap.curry0_apply, iteratedFDerivWithin_zero_apply] · intro m hm x hx - rcases (h x hx) m.succ (Order.add_one_le_of_lt hm) with ⟨u, hu, p, Hp⟩ + have : m < n := mod_cast hm + rcases (h x hx) m.succ (Order.add_one_le_of_lt this) with ⟨u, hu, p, Hp⟩ rw [insert_eq_of_mem hx] at hu rcases mem_nhdsWithin.1 hu with ⟨o, o_open, xo, ho⟩ rw [inter_comm] at ho @@ -533,15 +537,15 @@ protected theorem ContDiffOn.ftaylorSeriesWithin (h : ContDiffOn 𝕜 n f s) (hs change p y m = iteratedFDerivWithin 𝕜 m f s y rw [← iteratedFDerivWithin_inter_open o_open yo] exact - (Hp.mono ho).eq_iteratedFDerivWithin_of_uniqueDiffOn (WithTop.coe_le_coe.2 (Nat.le_succ m)) + (Hp.mono ho).eq_iteratedFDerivWithin_of_uniqueDiffOn (mod_cast Nat.le_succ m) (hs.inter o_open) ⟨hy, yo⟩ exact - ((Hp.mono ho).fderivWithin m (WithTop.coe_lt_coe.2 (lt_add_one m)) x ⟨hx, xo⟩).congr + ((Hp.mono ho).fderivWithin m (mod_cast lt_add_one m) x ⟨hx, xo⟩).congr (fun y hy => (A y hy).symm) (A x ⟨hx, xo⟩).symm · intro m hm apply continuousOn_of_locally_continuousOn intro x hx - rcases h x hx m hm with ⟨u, hu, p, Hp⟩ + rcases h x hx m (mod_cast hm) with ⟨u, hu, p, Hp⟩ rcases mem_nhdsWithin.1 hu with ⟨o, o_open, xo, ho⟩ rw [insert_eq_of_mem hx] at ho rw [inter_comm] at ho @@ -554,8 +558,8 @@ protected theorem ContDiffOn.ftaylorSeriesWithin (h : ContDiffOn 𝕜 n f s) (hs exact ((Hp.mono ho).cont m le_rfl).congr fun y hy => (A y hy).symm theorem contDiffOn_of_continuousOn_differentiableOn - (Hcont : ∀ m : ℕ, (m : ℕ∞) ≤ n → ContinuousOn (fun x => iteratedFDerivWithin 𝕜 m f s x) s) - (Hdiff : ∀ m : ℕ, (m : ℕ∞) < n → + (Hcont : ∀ m : ℕ, m ≤ n → ContinuousOn (fun x => iteratedFDerivWithin 𝕜 m f s x) s) + (Hdiff : ∀ m : ℕ, m < n → DifferentiableOn 𝕜 (fun x => iteratedFDerivWithin 𝕜 m f s x) s) : ContDiffOn 𝕜 n f s := by intro x hx m hm @@ -566,27 +570,27 @@ theorem contDiffOn_of_continuousOn_differentiableOn simp only [ftaylorSeriesWithin, ContinuousMultilinearMap.curry0_apply, iteratedFDerivWithin_zero_apply] · intro k hk y hy - convert (Hdiff k (lt_of_lt_of_le hk hm) y hy).hasFDerivWithinAt + convert (Hdiff k (lt_of_lt_of_le (mod_cast hk) hm) y hy).hasFDerivWithinAt · intro k hk - exact Hcont k (le_trans hk hm) + exact Hcont k (le_trans (mod_cast hk) hm) theorem contDiffOn_of_differentiableOn - (h : ∀ m : ℕ, (m : ℕ∞) ≤ n → DifferentiableOn 𝕜 (iteratedFDerivWithin 𝕜 m f s) s) : + (h : ∀ m : ℕ, m ≤ n → DifferentiableOn 𝕜 (iteratedFDerivWithin 𝕜 m f s) s) : ContDiffOn 𝕜 n f s := contDiffOn_of_continuousOn_differentiableOn (fun m hm => (h m hm).continuousOn) fun m hm => h m (le_of_lt hm) theorem ContDiffOn.continuousOn_iteratedFDerivWithin {m : ℕ} (h : ContDiffOn 𝕜 n f s) - (hmn : (m : ℕ∞) ≤ n) (hs : UniqueDiffOn 𝕜 s) : ContinuousOn (iteratedFDerivWithin 𝕜 m f s) s := - (h.ftaylorSeriesWithin hs).cont m hmn + (hmn : m ≤ n) (hs : UniqueDiffOn 𝕜 s) : ContinuousOn (iteratedFDerivWithin 𝕜 m f s) s := + (h.ftaylorSeriesWithin hs).cont m (mod_cast hmn) theorem ContDiffOn.differentiableOn_iteratedFDerivWithin {m : ℕ} (h : ContDiffOn 𝕜 n f s) - (hmn : (m : ℕ∞) < n) (hs : UniqueDiffOn 𝕜 s) : + (hmn : m < n) (hs : UniqueDiffOn 𝕜 s) : DifferentiableOn 𝕜 (iteratedFDerivWithin 𝕜 m f s) s := fun x hx => - ((h.ftaylorSeriesWithin hs).fderivWithin m hmn x hx).differentiableWithinAt + ((h.ftaylorSeriesWithin hs).fderivWithin m (mod_cast hmn) x hx).differentiableWithinAt theorem ContDiffWithinAt.differentiableWithinAt_iteratedFDerivWithin {m : ℕ} - (h : ContDiffWithinAt 𝕜 n f s x) (hmn : (m : ℕ∞) < n) (hs : UniqueDiffOn 𝕜 (insert x s)) : + (h : ContDiffWithinAt 𝕜 n f s x) (hmn : m < n) (hs : UniqueDiffOn 𝕜 (insert x s)) : DifferentiableWithinAt 𝕜 (iteratedFDerivWithin 𝕜 m f s) s x := by rcases h.contDiffOn' (Order.add_one_le_of_lt hmn) with ⟨u, uo, xu, hu⟩ set t := insert x s ∩ u @@ -605,14 +609,14 @@ theorem ContDiffWithinAt.differentiableWithinAt_iteratedFDerivWithin {m : ℕ} theorem contDiffOn_iff_continuousOn_differentiableOn (hs : UniqueDiffOn 𝕜 s) : ContDiffOn 𝕜 n f s ↔ - (∀ m : ℕ, (m : ℕ∞) ≤ n → ContinuousOn (fun x => iteratedFDerivWithin 𝕜 m f s x) s) ∧ - ∀ m : ℕ, (m : ℕ∞) < n → DifferentiableOn 𝕜 (fun x => iteratedFDerivWithin 𝕜 m f s x) s := + (∀ m : ℕ, m ≤ n → ContinuousOn (fun x => iteratedFDerivWithin 𝕜 m f s x) s) ∧ + ∀ m : ℕ, m < n → DifferentiableOn 𝕜 (fun x => iteratedFDerivWithin 𝕜 m f s x) s := ⟨fun h => ⟨fun _m hm => h.continuousOn_iteratedFDerivWithin hm hs, fun _m hm => h.differentiableOn_iteratedFDerivWithin hm hs⟩, fun h => contDiffOn_of_continuousOn_differentiableOn h.1 h.2⟩ theorem contDiffOn_succ_of_fderivWithin {n : ℕ} (hf : DifferentiableOn 𝕜 f s) - (h : ContDiffOn 𝕜 n (fun y => fderivWithin 𝕜 f s y) s) : ContDiffOn 𝕜 (n + 1 : ℕ) f s := by + (h : ContDiffOn 𝕜 n (fun y => fderivWithin 𝕜 f s y) s) : ContDiffOn 𝕜 (n + 1) f s := by intro x hx rw [contDiffWithinAt_succ_iff_hasFDerivWithinAt, insert_eq_of_mem hx] exact @@ -621,7 +625,7 @@ theorem contDiffOn_succ_of_fderivWithin {n : ℕ} (hf : DifferentiableOn 𝕜 f /-- A function is `C^(n + 1)` on a domain with unique derivatives if and only if it is differentiable there, and its derivative (expressed with `fderivWithin`) is `C^n`. -/ theorem contDiffOn_succ_iff_fderivWithin {n : ℕ} (hs : UniqueDiffOn 𝕜 s) : - ContDiffOn 𝕜 (n + 1 : ℕ) f s ↔ + ContDiffOn 𝕜 (n + 1) f s ↔ DifferentiableOn 𝕜 f s ∧ ContDiffOn 𝕜 n (fun y => fderivWithin 𝕜 f s y) s := by refine ⟨fun H => ?_, fun h => contDiffOn_succ_of_fderivWithin h.1 h.2⟩ refine ⟨H.differentiableOn (WithTop.coe_le_coe.2 (Nat.le_add_left 1 n)), fun x hx => ?_⟩ @@ -639,7 +643,7 @@ theorem contDiffOn_succ_iff_fderivWithin {n : ℕ} (hs : UniqueDiffOn 𝕜 s) : rwa [fderivWithin_inter (o_open.mem_nhds hy.2)] at A theorem contDiffOn_succ_iff_hasFDerivWithin {n : ℕ} (hs : UniqueDiffOn 𝕜 s) : - ContDiffOn 𝕜 (n + 1 : ℕ) f s ↔ + ContDiffOn 𝕜 (n + 1) f s ↔ ∃ f' : E → E →L[𝕜] F, ContDiffOn 𝕜 n f' s ∧ ∀ x, x ∈ s → HasFDerivWithinAt f (f' x) s x := by rw [contDiffOn_succ_iff_fderivWithin hs] refine ⟨fun h => ⟨fderivWithin 𝕜 f s, h.2, fun x hx => (h.1 x hx).hasFDerivWithinAt⟩, fun h => ?_⟩ @@ -650,7 +654,7 @@ theorem contDiffOn_succ_iff_hasFDerivWithin {n : ℕ} (hs : UniqueDiffOn 𝕜 s) /-- A function is `C^(n + 1)` on an open domain if and only if it is differentiable there, and its derivative (expressed with `fderiv`) is `C^n`. -/ theorem contDiffOn_succ_iff_fderiv_of_isOpen {n : ℕ} (hs : IsOpen s) : - ContDiffOn 𝕜 (n + 1 : ℕ) f s ↔ + ContDiffOn 𝕜 (n + 1) f s ↔ DifferentiableOn 𝕜 f s ∧ ContDiffOn 𝕜 n (fun y => fderiv 𝕜 f y) s := by rw [contDiffOn_succ_iff_fderivWithin hs.uniqueDiffOn] exact Iff.rfl.and (contDiffOn_congr fun x hx ↦ fderivWithin_of_isOpen hs hx) @@ -762,7 +766,7 @@ nonrec lemma ContDiffAt.contDiffOn {m : ℕ} (h : ContDiffAt 𝕜 n f x) (hm : m /-- A function is `C^(n + 1)` at a point iff locally, it has a derivative which is `C^n`. -/ theorem contDiffAt_succ_iff_hasFDerivAt {n : ℕ} : - ContDiffAt 𝕜 (n + 1 : ℕ) f x ↔ + ContDiffAt 𝕜 (n + 1) f x ↔ ∃ f' : E → E →L[𝕜] F, (∃ u ∈ 𝓝 x, ∀ x ∈ u, HasFDerivAt f (f' x) x) ∧ ContDiffAt 𝕜 n f' x := by rw [← contDiffWithinAt_univ, contDiffWithinAt_succ_iff_hasFDerivWithinAt] simp only [nhdsWithin_univ, exists_prop, mem_univ, insert_eq_of_mem] @@ -808,7 +812,8 @@ theorem contDiffOn_univ : ContDiffOn 𝕜 n f univ ↔ ContDiff 𝕜 n f := by rw [← hasFTaylorSeriesUpToOn_univ_iff] exact H.ftaylorSeriesWithin uniqueDiffOn_univ · rintro ⟨p, hp⟩ x _ m hm - exact ⟨univ, Filter.univ_sets _, p, (hp.hasFTaylorSeriesUpToOn univ).of_le hm⟩ + exact ⟨univ, Filter.univ_sets _, p, (hp.hasFTaylorSeriesUpToOn univ).of_le + (mod_cast hm)⟩ theorem contDiff_iff_contDiffAt : ContDiff 𝕜 n f ↔ ∀ x, ContDiffAt 𝕜 n f x := by simp [← contDiffOn_univ, ContDiffOn, ContDiffAt] @@ -839,8 +844,8 @@ theorem contDiffAt_zero : ContDiffAt 𝕜 0 f x ↔ ∃ u ∈ 𝓝 x, Continuous theorem contDiffAt_one_iff : ContDiffAt 𝕜 1 f x ↔ ∃ f' : E → E →L[𝕜] F, ∃ u ∈ 𝓝 x, ContinuousOn f' u ∧ ∀ x ∈ u, HasFDerivAt f (f' x) x := by - simp_rw [show (1 : ℕ∞) = (0 + 1 : ℕ) from (zero_add 1).symm, contDiffAt_succ_iff_hasFDerivAt, - show ((0 : ℕ) : ℕ∞) = 0 from rfl, contDiffAt_zero, + rw [show (1 : ℕ∞) = (0 : ℕ) + 1 from rfl, contDiffAt_succ_iff_hasFDerivAt] + simp_rw [show ((0 : ℕ) : ℕ∞) = 0 from rfl, contDiffAt_zero, exists_mem_and_iff antitone_bforall antitone_continuousOn, and_comm] theorem ContDiff.of_le (h : ContDiff 𝕜 n f) (hmn : m ≤ n) : ContDiff 𝕜 m f := @@ -864,7 +869,7 @@ theorem contDiff_iff_forall_nat_le : ContDiff 𝕜 n f ↔ ∀ m : ℕ, ↑m ≤ /-- A function is `C^(n+1)` iff it has a `C^n` derivative. -/ theorem contDiff_succ_iff_hasFDerivAt {n : ℕ} : - ContDiff 𝕜 (n + 1 : ℕ) f ↔ + ContDiff 𝕜 (n + 1) f ↔ ∃ f' : E → E →L[𝕜] F, ContDiff 𝕜 n f' ∧ ∀ x, HasFDerivAt f (f' x) x := by simp only [← contDiffOn_univ, ← hasFDerivWithinAt_univ, contDiffOn_succ_iff_hasFDerivWithin uniqueDiffOn_univ, Set.mem_univ, forall_true_left] @@ -884,30 +889,30 @@ theorem contDiff_iff_ftaylorSeries : theorem contDiff_iff_continuous_differentiable : ContDiff 𝕜 n f ↔ - (∀ m : ℕ, (m : ℕ∞) ≤ n → Continuous fun x => iteratedFDeriv 𝕜 m f x) ∧ - ∀ m : ℕ, (m : ℕ∞) < n → Differentiable 𝕜 fun x => iteratedFDeriv 𝕜 m f x := by + (∀ m : ℕ, m ≤ n → Continuous fun x => iteratedFDeriv 𝕜 m f x) ∧ + ∀ m : ℕ, m < n → Differentiable 𝕜 fun x => iteratedFDeriv 𝕜 m f x := by simp [contDiffOn_univ.symm, continuous_iff_continuousOn_univ, differentiableOn_univ.symm, iteratedFDerivWithin_univ, contDiffOn_iff_continuousOn_differentiableOn uniqueDiffOn_univ] /-- If `f` is `C^n` then its `m`-times iterated derivative is continuous for `m ≤ n`. -/ -theorem ContDiff.continuous_iteratedFDeriv {m : ℕ} (hm : (m : ℕ∞) ≤ n) (hf : ContDiff 𝕜 n f) : +theorem ContDiff.continuous_iteratedFDeriv {m : ℕ} (hm : m ≤ n) (hf : ContDiff 𝕜 n f) : Continuous fun x => iteratedFDeriv 𝕜 m f x := (contDiff_iff_continuous_differentiable.mp hf).1 m hm /-- If `f` is `C^n` then its `m`-times iterated derivative is differentiable for `m < n`. -/ -theorem ContDiff.differentiable_iteratedFDeriv {m : ℕ} (hm : (m : ℕ∞) < n) (hf : ContDiff 𝕜 n f) : +theorem ContDiff.differentiable_iteratedFDeriv {m : ℕ} (hm : m < n) (hf : ContDiff 𝕜 n f) : Differentiable 𝕜 fun x => iteratedFDeriv 𝕜 m f x := (contDiff_iff_continuous_differentiable.mp hf).2 m hm theorem contDiff_of_differentiable_iteratedFDeriv - (h : ∀ m : ℕ, (m : ℕ∞) ≤ n → Differentiable 𝕜 (iteratedFDeriv 𝕜 m f)) : ContDiff 𝕜 n f := + (h : ∀ m : ℕ, m ≤ n → Differentiable 𝕜 (iteratedFDeriv 𝕜 m f)) : ContDiff 𝕜 n f := contDiff_iff_continuous_differentiable.2 ⟨fun m hm => (h m hm).continuous, fun m hm => h m (le_of_lt hm)⟩ /-- A function is `C^(n + 1)` if and only if it is differentiable, and its derivative (formulated in terms of `fderiv`) is `C^n`. -/ theorem contDiff_succ_iff_fderiv {n : ℕ} : - ContDiff 𝕜 (n + 1 : ℕ) f ↔ Differentiable 𝕜 f ∧ ContDiff 𝕜 n fun y => fderiv 𝕜 f y := by + ContDiff 𝕜 (n + 1) f ↔ Differentiable 𝕜 f ∧ ContDiff 𝕜 n fun y => fderiv 𝕜 f y := by simp only [← contDiffOn_univ, ← differentiableOn_univ, ← fderivWithin_univ, contDiffOn_succ_iff_fderivWithin uniqueDiffOn_univ] diff --git a/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean b/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean index 149c86f2c306e..abdac0002a6e2 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean @@ -99,7 +99,7 @@ In this file, we denote `⊤ : ℕ∞` with `∞`. noncomputable section open scoped Classical -open NNReal Topology Filter +open ENat NNReal Topology Filter local notation "∞" => (⊤ : ℕ∞) @@ -115,7 +115,8 @@ universe u uE uF variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] {E : Type uE} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {F : Type uF} [NormedAddCommGroup F] [NormedSpace 𝕜 F] - {s t u : Set E} {f f₁ : E → F} {x : E} {m n : ℕ∞} {p : E → FormalMultilinearSeries 𝕜 E F} + {s t u : Set E} {f f₁ : E → F} {x : E} {m n N : WithTop ℕ∞} + {p : E → FormalMultilinearSeries 𝕜 E F} /-! ### Functions with a Taylor series on a domain -/ @@ -125,12 +126,12 @@ derivative of `p m` for `m < n`, and is continuous for `m ≤ n`. This is a pred Notice that `p` does not sum up to `f` on the diagonal (`FormalMultilinearSeries.sum`), even if `f` is analytic and `n = ∞`: an additional `1/m!` factor on the `m`th term is necessary for that. -/ -structure HasFTaylorSeriesUpToOn (n : ℕ∞) (f : E → F) (p : E → FormalMultilinearSeries 𝕜 E F) - (s : Set E) : Prop where +structure HasFTaylorSeriesUpToOn + (n : WithTop ℕ∞) (f : E → F) (p : E → FormalMultilinearSeries 𝕜 E F) (s : Set E) : Prop where zero_eq : ∀ x ∈ s, (p x 0).curry0 = f x protected fderivWithin : ∀ m : ℕ, (m : ℕ∞) < n → ∀ x ∈ s, HasFDerivWithinAt (p · m) (p x m.succ).curryLeft s x - cont : ∀ m : ℕ, (m : ℕ∞) ≤ n → ContinuousOn (p · m) s + cont : ∀ m : ℕ, m ≤ n → ContinuousOn (p · m) s theorem HasFTaylorSeriesUpToOn.zero_eq' (h : HasFTaylorSeriesUpToOn n f p s) {x : E} (hx : x ∈ s) : p x 0 = (continuousMultilinearCurryFin0 𝕜 E F).symm (f x) := by @@ -170,26 +171,31 @@ theorem hasFTaylorSeriesUpToOn_zero_iff : rw [continuousOn_congr this, LinearIsometryEquiv.comp_continuousOn_iff] exact H.1 -theorem hasFTaylorSeriesUpToOn_top_iff : - HasFTaylorSeriesUpToOn ∞ f p s ↔ ∀ n : ℕ, HasFTaylorSeriesUpToOn n f p s := by +theorem hasFTaylorSeriesUpToOn_top_iff_add (hN : ∞ ≤ N) (k : ℕ) : + HasFTaylorSeriesUpToOn N f p s ↔ ∀ n : ℕ, HasFTaylorSeriesUpToOn (n + k : ℕ) f p s := by constructor - · intro H n; exact H.of_le le_top + · intro H n + apply H.of_le (natCast_le_of_coe_top_le_withTop hN _) · intro H constructor · exact (H 0).zero_eq · intro m _ - apply (H m.succ).fderivWithin m (WithTop.coe_lt_coe.2 (lt_add_one m)) + apply (H m.succ).fderivWithin m (by norm_cast; omega) · intro m _ - apply (H m).cont m le_rfl + apply (H m).cont m (by simp) + +theorem hasFTaylorSeriesUpToOn_top_iff (hN : ∞ ≤ N) : + HasFTaylorSeriesUpToOn N f p s ↔ ∀ n : ℕ, HasFTaylorSeriesUpToOn n f p s := by + simpa using hasFTaylorSeriesUpToOn_top_iff_add hN 0 /-- In the case that `n = ∞` we don't need the continuity assumption in `HasFTaylorSeriesUpToOn`. -/ -theorem hasFTaylorSeriesUpToOn_top_iff' : - HasFTaylorSeriesUpToOn ∞ f p s ↔ +theorem hasFTaylorSeriesUpToOn_top_iff' (hN : ∞ ≤ N) : + HasFTaylorSeriesUpToOn N f p s ↔ (∀ x ∈ s, (p x 0).curry0 = f x) ∧ - ∀ m : ℕ, ∀ x ∈ s, HasFDerivWithinAt (fun y => p y m) (p x m.succ).curryLeft s x := + ∀ m : ℕ, ∀ x ∈ s, HasFDerivWithinAt (fun y => p y m) (p x m.succ).curryLeft s x := by -- Everything except for the continuity is trivial: - ⟨fun h => ⟨h.1, fun m => h.2 m (WithTop.coe_lt_top m)⟩, fun h => + refine ⟨fun h => ⟨h.1, fun m => h.2 m (natCast_lt_of_coe_top_le_withTop hN _)⟩, fun h => ⟨h.1, fun m _ => h.2 m, fun m _ x hx => -- The continuity follows from the existence of a derivative: (h.2 m x hx).continuousWithinAt⟩⟩ @@ -241,21 +247,21 @@ theorem hasFTaylorSeriesUpToOn_succ_iff_left {n : ℕ} : (∀ x ∈ s, HasFDerivWithinAt (fun y => p y n) (p x n.succ).curryLeft s x) ∧ ContinuousOn (fun x => p x (n + 1)) s := by constructor - · exact fun h ↦ ⟨h.of_le (WithTop.coe_le_coe.2 (Nat.le_succ n)), - h.fderivWithin _ (WithTop.coe_lt_coe.2 (lt_add_one n)), h.cont (n + 1) le_rfl⟩ + · exact fun h ↦ ⟨h.of_le (mod_cast Nat.le_succ n), + h.fderivWithin _ (mod_cast lt_add_one n), h.cont (n + 1) le_rfl⟩ · intro h constructor · exact h.1.zero_eq · intro m hm by_cases h' : m < n - · exact h.1.fderivWithin m (WithTop.coe_lt_coe.2 h') - · have : m = n := Nat.eq_of_lt_succ_of_not_lt (WithTop.coe_lt_coe.1 hm) h' + · exact h.1.fderivWithin m (mod_cast h') + · have : m = n := Nat.eq_of_lt_succ_of_not_lt (mod_cast hm) h' rw [this] exact h.2.1 · intro m hm by_cases h' : m ≤ n - · apply h.1.cont m (WithTop.coe_le_coe.2 h') - · have : m = n + 1 := le_antisymm (WithTop.coe_le_coe.1 hm) (not_le.1 h') + · apply h.1.cont m (mod_cast h') + · have : m = n + 1 := le_antisymm (mod_cast hm) (not_le.1 h') rw [this] exact h.2.2 @@ -273,8 +279,8 @@ theorem HasFTaylorSeriesUpToOn.shift_of_succ constructor · intro x _ rfl - · intro m (hm : (m : ℕ∞) < n) x (hx : x ∈ s) - have A : (m.succ : ℕ∞) < n.succ := by + · intro m (hm : (m : WithTop ℕ∞) < n) x (hx : x ∈ s) + have A : (m.succ : WithTop ℕ∞) < n.succ := by rw [Nat.cast_lt] at hm ⊢ exact Nat.succ_lt_succ hm change HasFDerivWithinAt (continuousMultilinearCurryRightEquiv' 𝕜 m E F ∘ (p · m.succ)) @@ -284,7 +290,7 @@ theorem HasFTaylorSeriesUpToOn.shift_of_succ ext y v change p x (m + 2) (snoc (cons y (init v)) (v (last _))) = p x (m + 2) (cons y v) rw [← cons_snoc_eq_snoc_cons, snoc_init_self] - · intro m (hm : (m : ℕ∞) ≤ n) + · intro m (hm : (m : WithTop ℕ∞) ≤ n) suffices A : ContinuousOn (p · (m + 1)) s from (continuousMultilinearCurryRightEquiv' 𝕜 m E F).continuous.comp_continuousOn A refine H.cont _ ?_ @@ -292,8 +298,8 @@ theorem HasFTaylorSeriesUpToOn.shift_of_succ exact Nat.succ_le_succ hm /-- `p` is a Taylor series of `f` up to `n+1` if and only if `p.shift` is a Taylor series up to `n` -for `p 1`, which is a derivative of `f`. -/ -theorem hasFTaylorSeriesUpToOn_succ_iff_right {n : ℕ} : +for `p 1`, which is a derivative of `f`. Version for `n : ℕ`. -/ +theorem hasFTaylorSeriesUpToOn_succ_nat_iff_right {n : ℕ} : HasFTaylorSeriesUpToOn (n + 1 : ℕ) f p s ↔ (∀ x ∈ s, (p x 0).curry0 = f x) ∧ (∀ x ∈ s, HasFDerivWithinAt (fun y => p y 0) (p x 1).curryLeft s x) ∧ @@ -306,10 +312,10 @@ theorem hasFTaylorSeriesUpToOn_succ_iff_right {n : ℕ} : · rintro ⟨Hzero_eq, Hfderiv_zero, Htaylor⟩ constructor · exact Hzero_eq - · intro m (hm : (m : ℕ∞) < n.succ) x (hx : x ∈ s) + · intro m (hm : (m : WithTop ℕ∞) < n.succ) x (hx : x ∈ s) cases' m with m · exact Hfderiv_zero x hx - · have A : (m : ℕ∞) < n := by + · have A : (m : WithTop ℕ∞) < n := by rw [Nat.cast_lt] at hm ⊢ exact Nat.lt_of_succ_lt_succ hm have : @@ -322,7 +328,7 @@ theorem hasFTaylorSeriesUpToOn_succ_iff_right {n : ℕ} : (p x (Nat.succ (Nat.succ m))) (cons y v) = (p x m.succ.succ) (snoc (cons y (init v)) (v (last _))) rw [← cons_snoc_eq_snoc_cons, snoc_init_self] - · intro m (hm : (m : ℕ∞) ≤ n.succ) + · intro m (hm : (m : WithTop ℕ∞) ≤ n.succ) cases' m with m · have : DifferentiableOn 𝕜 (fun x => p x 0) s := fun x hx => (Hfderiv_zero x hx).differentiableWithinAt @@ -332,6 +338,37 @@ theorem hasFTaylorSeriesUpToOn_succ_iff_right {n : ℕ} : rw [Nat.cast_le] at hm ⊢ exact Nat.lt_succ_iff.mp hm +/-- `p` is a Taylor series of `f` up to `⊤` if and only if `p.shift` is a Taylor series up to `⊤` +for `p 1`, which is a derivative of `f`. -/ +theorem hasFTaylorSeriesUpToOn_top_iff_right (hN : ∞ ≤ N) : + HasFTaylorSeriesUpToOn N f p s ↔ + (∀ x ∈ s, (p x 0).curry0 = f x) ∧ + (∀ x ∈ s, HasFDerivWithinAt (fun y => p y 0) (p x 1).curryLeft s x) ∧ + HasFTaylorSeriesUpToOn N (fun x => continuousMultilinearCurryFin1 𝕜 E F (p x 1)) + (fun x => (p x).shift) s := by + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · rw [hasFTaylorSeriesUpToOn_top_iff_add hN 1] at h + rw [hasFTaylorSeriesUpToOn_top_iff hN] + exact ⟨(hasFTaylorSeriesUpToOn_succ_nat_iff_right.1 (h 1)).1, + (hasFTaylorSeriesUpToOn_succ_nat_iff_right.1 (h 1)).2.1, + fun n ↦ (hasFTaylorSeriesUpToOn_succ_nat_iff_right.1 (h n)).2.2⟩ + · apply (hasFTaylorSeriesUpToOn_top_iff_add hN 1).2 (fun n ↦ ?_) + rw [hasFTaylorSeriesUpToOn_succ_nat_iff_right] + exact ⟨h.1, h.2.1, (h.2.2).of_le (m := n) (natCast_le_of_coe_top_le_withTop hN n)⟩ + +/-- `p` is a Taylor series of `f` up to `n+1` if and only if `p.shift` is a Taylor series up to `n` +for `p 1`, which is a derivative of `f`. Version for `n : WithTop ℕ∞`. -/ +theorem hasFTaylorSeriesUpToOn_succ_iff_right : + HasFTaylorSeriesUpToOn (n + 1) f p s ↔ + (∀ x ∈ s, (p x 0).curry0 = f x) ∧ + (∀ x ∈ s, HasFDerivWithinAt (fun y => p y 0) (p x 1).curryLeft s x) ∧ + HasFTaylorSeriesUpToOn n (fun x => continuousMultilinearCurryFin1 𝕜 E F (p x 1)) + (fun x => (p x).shift) s := by + match n with + | ⊤ => exact hasFTaylorSeriesUpToOn_top_iff_right (by simp) + | (⊤ : ℕ∞) => exact hasFTaylorSeriesUpToOn_top_iff_right (by simp) + | (n : ℕ) => exact hasFTaylorSeriesUpToOn_succ_nat_iff_right + /-! ### Iterated derivative within a set -/ @@ -537,11 +574,11 @@ theorem iteratedFDerivWithin_inter_open {n : ℕ} (hu : IsOpen u) (hx : x ∈ u) /-- On a set with unique differentiability, any choice of iterated differential has to coincide with the one we have chosen in `iteratedFDerivWithin 𝕜 m f s`. -/ theorem HasFTaylorSeriesUpToOn.eq_iteratedFDerivWithin_of_uniqueDiffOn - (h : HasFTaylorSeriesUpToOn n f p s) {m : ℕ} (hmn : (m : ℕ∞) ≤ n) (hs : UniqueDiffOn 𝕜 s) + (h : HasFTaylorSeriesUpToOn n f p s) {m : ℕ} (hmn : m ≤ n) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) : p x m = iteratedFDerivWithin 𝕜 m f s x := by induction' m with m IH generalizing x · rw [h.zero_eq' hx, iteratedFDerivWithin_zero_eq_comp]; rfl - · have A : (m : ℕ∞) < n := lt_of_lt_of_le (WithTop.coe_lt_coe.2 (lt_add_one m)) hmn + · have A : (m : ℕ∞) < n := lt_of_lt_of_le (mod_cast lt_add_one m) hmn have : HasFDerivWithinAt (fun y : E => iteratedFDerivWithin 𝕜 m f s y) (ContinuousMultilinearMap.curryLeft (p x (Nat.succ m))) s x := @@ -563,11 +600,11 @@ derivative of `p m` for `m < n`, and is continuous for `m ≤ n`. This is a pred Notice that `p` does not sum up to `f` on the diagonal (`FormalMultilinearSeries.sum`), even if `f` is analytic and `n = ∞`: an addition `1/m!` factor on the `m`th term is necessary for that. -/ -structure HasFTaylorSeriesUpTo (n : ℕ∞) (f : E → F) (p : E → FormalMultilinearSeries 𝕜 E F) : - Prop where +structure HasFTaylorSeriesUpTo + (n : WithTop ℕ∞) (f : E → F) (p : E → FormalMultilinearSeries 𝕜 E F) : Prop where zero_eq : ∀ x, (p x 0).curry0 = f x - fderiv : ∀ m : ℕ, (m : ℕ∞) < n → ∀ x, HasFDerivAt (fun y => p y m) (p x m.succ).curryLeft x - cont : ∀ m : ℕ, (m : ℕ∞) ≤ n → Continuous fun x => p x m + fderiv : ∀ m : ℕ, m < n → ∀ x, HasFDerivAt (fun y => p y m) (p x m.succ).curryLeft x + cont : ∀ m : ℕ, m ≤ n → Continuous fun x => p x m theorem HasFTaylorSeriesUpTo.zero_eq' (h : HasFTaylorSeriesUpTo n f p) (x : E) : p x 0 = (continuousMultilinearCurryFin0 𝕜 E F).symm (f x) := by @@ -600,10 +637,13 @@ theorem HasFTaylorSeriesUpTo.hasFTaylorSeriesUpToOn (h : HasFTaylorSeriesUpTo n HasFTaylorSeriesUpToOn n f p s := (hasFTaylorSeriesUpToOn_univ_iff.2 h).mono (subset_univ _) -theorem HasFTaylorSeriesUpTo.ofLe (h : HasFTaylorSeriesUpTo n f p) (hmn : m ≤ n) : +theorem HasFTaylorSeriesUpTo.of_le (h : HasFTaylorSeriesUpTo n f p) (hmn : m ≤ n) : HasFTaylorSeriesUpTo m f p := by rw [← hasFTaylorSeriesUpToOn_univ_iff] at h ⊢; exact h.of_le hmn +@[deprecated (since := "2024-11-07")] +alias HasFTaylorSeriesUpTo.ofLe := HasFTaylorSeriesUpTo.of_le + theorem HasFTaylorSeriesUpTo.continuous (h : HasFTaylorSeriesUpTo n f p) : Continuous f := by rw [← hasFTaylorSeriesUpToOn_univ_iff] at h rw [continuous_iff_continuousOn_univ] @@ -614,17 +654,17 @@ theorem hasFTaylorSeriesUpTo_zero_iff : simp [hasFTaylorSeriesUpToOn_univ_iff.symm, continuous_iff_continuousOn_univ, hasFTaylorSeriesUpToOn_zero_iff] -theorem hasFTaylorSeriesUpTo_top_iff : - HasFTaylorSeriesUpTo ∞ f p ↔ ∀ n : ℕ, HasFTaylorSeriesUpTo n f p := by - simp only [← hasFTaylorSeriesUpToOn_univ_iff, hasFTaylorSeriesUpToOn_top_iff] +theorem hasFTaylorSeriesUpTo_top_iff (hN : ∞ ≤ N) : + HasFTaylorSeriesUpTo N f p ↔ ∀ n : ℕ, HasFTaylorSeriesUpTo n f p := by + simp only [← hasFTaylorSeriesUpToOn_univ_iff, hasFTaylorSeriesUpToOn_top_iff hN] /-- In the case that `n = ∞` we don't need the continuity assumption in `HasFTaylorSeriesUpTo`. -/ -theorem hasFTaylorSeriesUpTo_top_iff' : - HasFTaylorSeriesUpTo ∞ f p ↔ +theorem hasFTaylorSeriesUpTo_top_iff' (hN : ∞ ≤ N) : + HasFTaylorSeriesUpTo N f p ↔ (∀ x, (p x 0).curry0 = f x) ∧ ∀ (m : ℕ) (x), HasFDerivAt (fun y => p y m) (p x m.succ).curryLeft x := by - simp only [← hasFTaylorSeriesUpToOn_univ_iff, hasFTaylorSeriesUpToOn_top_iff', mem_univ, + simp only [← hasFTaylorSeriesUpToOn_univ_iff, hasFTaylorSeriesUpToOn_top_iff' hN, mem_univ, forall_true_left, hasFDerivWithinAt_univ] /-- If a function has a Taylor series at order at least `1`, then the term of order `1` of this @@ -639,15 +679,17 @@ theorem HasFTaylorSeriesUpTo.differentiable (h : HasFTaylorSeriesUpTo n f p) (hn /-- `p` is a Taylor series of `f` up to `n+1` if and only if `p.shift` is a Taylor series up to `n` for `p 1`, which is a derivative of `f`. -/ -theorem hasFTaylorSeriesUpTo_succ_iff_right {n : ℕ} : +theorem hasFTaylorSeriesUpTo_succ_nat_iff_right {n : ℕ} : HasFTaylorSeriesUpTo (n + 1 : ℕ) f p ↔ (∀ x, (p x 0).curry0 = f x) ∧ (∀ x, HasFDerivAt (fun y => p y 0) (p x 1).curryLeft x) ∧ HasFTaylorSeriesUpTo n (fun x => continuousMultilinearCurryFin1 𝕜 E F (p x 1)) fun x => (p x).shift := by - simp only [hasFTaylorSeriesUpToOn_succ_iff_right, ← hasFTaylorSeriesUpToOn_univ_iff, mem_univ, + simp only [hasFTaylorSeriesUpToOn_succ_nat_iff_right, ← hasFTaylorSeriesUpToOn_univ_iff, mem_univ, forall_true_left, hasFDerivWithinAt_univ] +@[deprecated (since := "2024-11-07")] +alias hasFTaylorSeriesUpTo_succ_iff_right := hasFTaylorSeriesUpTo_succ_nat_iff_right /-! ### Iterated derivative -/ diff --git a/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean b/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean index 1c2e5249d4f46..582c9ca4fec35 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean @@ -917,7 +917,7 @@ private lemma faaDiBruno_aux2 {m : ℕ} (q : FormalMultilinearSeries 𝕜 F G) /-- *Faa di Bruno* formula: If two functions `g` and `f` have Taylor series up to `n` given by `q` and `p`, then `g ∘ f` also has a Taylor series, given by `q.taylorComp p`. -/ -theorem HasFTaylorSeriesUpToOn.comp {n : ℕ∞} {g : F → G} {f : E → F} +theorem HasFTaylorSeriesUpToOn.comp {n : WithTop ℕ∞} {g : F → G} {f : E → F} (hg : HasFTaylorSeriesUpToOn n g q t) (hf : HasFTaylorSeriesUpToOn n f p s) (h : MapsTo f s t) : HasFTaylorSeriesUpToOn n (g ∘ f) (fun x ↦ (q (f x)).taylorComp (p x)) s := by /- One has to check that the `m+1`-th term is the derivative of the `m`-th term. The `m`-th term @@ -940,8 +940,8 @@ theorem HasFTaylorSeriesUpToOn.comp {n : ℕ∞} {g : F → G} {f : E → F} change HasFDerivWithinAt (fun y ↦ B (q (f y) c.length) (fun i ↦ p y (c.partSize i))) (∑ i : Option (Fin c.length), ((q (f x)).compAlongOrderedFinpartition (p x) (c.extend i)).curryLeft) s x - have cm : (c.length : ℕ∞) ≤ m := by exact_mod_cast OrderedFinpartition.length_le c - have cp i : (c.partSize i : ℕ∞) ≤ m := by + have cm : (c.length : WithTop ℕ∞) ≤ m := mod_cast OrderedFinpartition.length_le c + have cp i : (c.partSize i : WithTop ℕ∞) ≤ m := by exact_mod_cast OrderedFinpartition.partSize_le c i have I i : HasFDerivWithinAt (fun x ↦ p x (c.partSize i)) (p x (c.partSize i).succ).curryLeft s x := @@ -949,7 +949,8 @@ theorem HasFTaylorSeriesUpToOn.comp {n : ℕ∞} {g : F → G} {f : E → F} have J : HasFDerivWithinAt (fun x ↦ q x c.length) (q (f x) c.length.succ).curryLeft t (f x) := hg.fderivWithin c.length (cm.trans_lt hm) (f x) (h hx) have K : HasFDerivWithinAt f ((continuousMultilinearCurryFin1 𝕜 E F) (p x 1)) s x := - hf.hasFDerivWithinAt (le_trans le_add_self (Order.add_one_le_of_lt hm)) hx + hf.hasFDerivWithinAt (le_trans (mod_cast Nat.le_add_left 1 m) + (ENat.add_one_natCast_le_withTop_of_lt hm)) hx convert HasFDerivWithinAt.linear_multilinear_comp (J.comp x K h) I B simp only [Nat.succ_eq_add_one, Fintype.sum_option, comp_apply, faaDiBruno_aux1, faaDiBruno_aux2] @@ -975,8 +976,9 @@ theorem HasFTaylorSeriesUpToOn.comp {n : ℕ∞} {g : F → G} {f : E → F} change ContinuousOn ((fun p ↦ B p.1 p.2) ∘ (fun x ↦ (q (f x) c.length, fun i ↦ p x (c.partSize i)))) s apply B.continuous_uncurry_of_multilinear.comp_continuousOn (ContinuousOn.prod ?_ ?_) - · have : (c.length : ℕ∞) ≤ m := by exact_mod_cast OrderedFinpartition.length_le c + · have : (c.length : WithTop ℕ∞) ≤ m := mod_cast OrderedFinpartition.length_le c exact (hg.cont c.length (this.trans hm)).comp hf.continuousOn h · apply continuousOn_pi.2 (fun i ↦ ?_) - have : (c.partSize i : ℕ∞) ≤ m := by exact_mod_cast OrderedFinpartition.partSize_le c i + have : (c.partSize i : WithTop ℕ∞) ≤ m := by + exact_mod_cast OrderedFinpartition.partSize_le c i exact hf.cont _ (this.trans hm) diff --git a/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean b/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean index e5d673c7ffac0..555d81bb107e4 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean @@ -54,17 +54,17 @@ often requires an inconvenient need to generalize `F`, which results in universe This lemma avoids these universe issues, but only applies for finite dimensional `E`. -/ theorem contDiff_succ_iff_fderiv_apply [FiniteDimensional 𝕜 D] {n : ℕ} {f : D → E} : - ContDiff 𝕜 (n + 1 : ℕ) f ↔ Differentiable 𝕜 f ∧ ∀ y, ContDiff 𝕜 n fun x => fderiv 𝕜 f x y := by + ContDiff 𝕜 (n + 1) f ↔ Differentiable 𝕜 f ∧ ∀ y, ContDiff 𝕜 n fun x => fderiv 𝕜 f x y := by rw [contDiff_succ_iff_fderiv, contDiff_clm_apply_iff] theorem contDiffOn_succ_of_fderiv_apply [FiniteDimensional 𝕜 D] {n : ℕ} {f : D → E} {s : Set D} (hf : DifferentiableOn 𝕜 f s) (h : ∀ y, ContDiffOn 𝕜 n (fun x => fderivWithin 𝕜 f s x y) s) : - ContDiffOn 𝕜 (n + 1 : ℕ) f s := + ContDiffOn 𝕜 (n + 1) f s := contDiffOn_succ_of_fderivWithin hf <| contDiffOn_clm_apply.mpr h theorem contDiffOn_succ_iff_fderiv_apply [FiniteDimensional 𝕜 D] {n : ℕ} {f : D → E} {s : Set D} (hs : UniqueDiffOn 𝕜 s) : - ContDiffOn 𝕜 (n + 1 : ℕ) f s ↔ + ContDiffOn 𝕜 (n + 1) f s ↔ DifferentiableOn 𝕜 f s ∧ ∀ y, ContDiffOn 𝕜 n (fun x => fderivWithin 𝕜 f s x y) s := by rw [contDiffOn_succ_iff_fderivWithin hs, contDiffOn_clm_apply] diff --git a/Mathlib/Analysis/Calculus/ContDiff/RCLike.lean b/Mathlib/Analysis/Calculus/ContDiff/RCLike.lean index d9f341edf671a..5f8fdba324bcb 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/RCLike.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/RCLike.lean @@ -29,7 +29,8 @@ variable {n : ℕ∞} {𝕂 : Type*} [RCLike 𝕂] {E' : Type*} [NormedAddCommGr /-- If a function has a Taylor series at order at least 1, then at points in the interior of the domain of definition, the term of order 1 of this series is a strict derivative of `f`. -/ -theorem HasFTaylorSeriesUpToOn.hasStrictFDerivAt {s : Set E'} {f : E' → F'} {x : E'} +theorem HasFTaylorSeriesUpToOn.hasStrictFDerivAt {n : WithTop ℕ∞} + {s : Set E'} {f : E' → F'} {x : E'} {p : E' → FormalMultilinearSeries 𝕂 E' F'} (hf : HasFTaylorSeriesUpToOn n f p s) (hn : 1 ≤ n) (hs : s ∈ 𝓝 x) : HasStrictFDerivAt f ((continuousMultilinearCurryFin1 𝕂 E' F') (p x 1)) x := hasStrictFDerivAt_of_hasFDerivAt_of_continuousAt (hf.eventually_hasFDerivAt hn hs) <| diff --git a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean index 76c137c4bab78..055c8cb25720e 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean @@ -382,7 +382,7 @@ protected theorem AnalyticOn.iteratedFDerivWithin (h : AnalyticOn 𝕜 f s) apply AnalyticOnNhd.comp_analyticOn _ (IH.fderivWithin hu) (mapsTo_univ _ _) apply LinearIsometryEquiv.analyticOnNhd -lemma AnalyticOn.hasFTaylorSeriesUpToOn {n : ℕ∞} +lemma AnalyticOn.hasFTaylorSeriesUpToOn {n : WithTop ℕ∞} (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 ↦ ?_⟩ diff --git a/Mathlib/Analysis/Convolution.lean b/Mathlib/Analysis/Convolution.lean index a50566576f0f1..e3a80a3cd256e 100644 --- a/Mathlib/Analysis/Convolution.lean +++ b/Mathlib/Analysis/Convolution.lean @@ -1187,7 +1187,8 @@ theorem contDiffOn_convolution_right_with_param_aux {G : Type uP} {E' : Type uP} have A : ∀ q₀ : P × G, q₀.1 ∈ s → HasFDerivAt (fun q : P × G => (f ⋆[L, μ] g q.1) q.2) (f' q₀.1 q₀.2) q₀ := hasFDerivAt_convolution_right_with_param L hs hk hgs hf hg.one_of_succ - rw [contDiffOn_succ_iff_fderiv_of_isOpen (hs.prod (@isOpen_univ G _))] at hg ⊢ + rw [show ((n + 1 : ℕ) : ℕ∞) = n + 1 from rfl, + contDiffOn_succ_iff_fderiv_of_isOpen (hs.prod (@isOpen_univ G _))] at hg ⊢ constructor · rintro ⟨p, x⟩ ⟨hp, -⟩ exact (A (p, x) hp).differentiableAt.differentiableWithinAt diff --git a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean index 7a53f32811c4c..5b882dad07252 100644 --- a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean +++ b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean @@ -440,7 +440,7 @@ lemma integrable_fourierPowSMulRight {n : ℕ} (hf : Integrable (fun v ↦ ‖v filter_upwards with v exact (norm_fourierPowSMulRight_le L f v n).trans (le_of_eq (by ring)) -lemma hasFTaylorSeriesUpTo_fourierIntegral {N : ℕ∞} +lemma hasFTaylorSeriesUpTo_fourierIntegral {N : WithTop ℕ∞} (hf : ∀ (n : ℕ), n ≤ N → Integrable (fun v ↦ ‖v‖^n * ‖f v‖) μ) (h'f : AEStronglyMeasurable f μ) : HasFTaylorSeriesUpTo N (fourierIntegral 𝐞 μ L.toLinearMap₂ f) @@ -455,7 +455,8 @@ lemma hasFTaylorSeriesUpTo_fourierIntegral {N : ℕ∞} have I₁ : Integrable (fun v ↦ fourierPowSMulRight L f v n) μ := integrable_fourierPowSMulRight L (hf n hn.le) h'f have I₂ : Integrable (fun v ↦ ‖v‖ * ‖fourierPowSMulRight L f v n‖) μ := by - apply ((hf (n+1) (Order.add_one_le_of_lt hn)).const_mul ((2 * π * ‖L‖) ^ n)).mono' + apply ((hf (n+1) (ENat.add_one_natCast_le_withTop_of_lt hn)).const_mul + ((2 * π * ‖L‖) ^ n)).mono' (continuous_norm.aestronglyMeasurable.mul (h'f.fourierPowSMulRight L n).norm) filter_upwards with v simp only [Pi.mul_apply, norm_mul, norm_norm] @@ -465,7 +466,7 @@ lemma hasFTaylorSeriesUpTo_fourierIntegral {N : ℕ∞} gcongr; apply norm_fourierPowSMulRight_le _ = (2 * π * ‖L‖) ^ n * (‖v‖ ^ (n + 1) * ‖f v‖) := by rw [pow_succ]; ring have I₃ : Integrable (fun v ↦ fourierPowSMulRight L f v (n + 1)) μ := - integrable_fourierPowSMulRight L (hf (n + 1) (Order.add_one_le_of_lt hn)) h'f + integrable_fourierPowSMulRight L (hf (n + 1) (ENat.add_one_natCast_le_withTop_of_lt hn)) h'f have I₄ : Integrable (fun v ↦ fourierSMulRight L (fun v ↦ fourierPowSMulRight L f v n) v) μ := by apply (I₂.const_mul ((2 * π * ‖L‖))).mono' (h'f.fourierPowSMulRight L n).fourierSMulRight @@ -488,12 +489,22 @@ lemma hasFTaylorSeriesUpTo_fourierIntegral {N : ℕ∞} apply fourierIntegral_continuous Real.continuous_fourierChar (by apply L.continuous₂) exact integrable_fourierPowSMulRight L (hf n hn) h'f +/-- Variant of `hasFTaylorSeriesUpTo_fourierIntegral` in which the smoothness index is restricted +to `ℕ∞` (and so are the inequalities in the assumption `hf`). Avoids normcasting in some +applications. -/ +lemma hasFTaylorSeriesUpTo_fourierIntegral' {N : ℕ∞} + (hf : ∀ (n : ℕ), n ≤ N → Integrable (fun v ↦ ‖v‖^n * ‖f v‖) μ) + (h'f : AEStronglyMeasurable f μ) : + HasFTaylorSeriesUpTo N (fourierIntegral 𝐞 μ L.toLinearMap₂ f) + (fun w n ↦ fourierIntegral 𝐞 μ L.toLinearMap₂ (fun v ↦ fourierPowSMulRight L f v n) w) := + hasFTaylorSeriesUpTo_fourierIntegral _ (fun n hn ↦ hf n (mod_cast hn)) h'f + /-- If `‖v‖^n * ‖f v‖` is integrable for all `n ≤ N`, then the Fourier transform of `f` is `C^N`. -/ theorem contDiff_fourierIntegral {N : ℕ∞} (hf : ∀ (n : ℕ), n ≤ N → Integrable (fun v ↦ ‖v‖^n * ‖f v‖) μ) : ContDiff ℝ N (fourierIntegral 𝐞 μ L.toLinearMap₂ f) := by by_cases h'f : Integrable f μ - · exact (hasFTaylorSeriesUpTo_fourierIntegral L hf h'f.1).contDiff + · exact (hasFTaylorSeriesUpTo_fourierIntegral' L hf h'f.1).contDiff · have : fourierIntegral 𝐞 μ L.toLinearMap₂ f = 0 := by ext w; simp [fourierIntegral, integral, h'f] simpa [this] using contDiff_const @@ -507,7 +518,8 @@ lemma iteratedFDeriv_fourierIntegral {N : ℕ∞} iteratedFDeriv ℝ n (fourierIntegral 𝐞 μ L.toLinearMap₂ f) = fourierIntegral 𝐞 μ L.toLinearMap₂ (fun v ↦ fourierPowSMulRight L f v n) := by ext w : 1 - exact ((hasFTaylorSeriesUpTo_fourierIntegral L hf h'f).eq_iteratedFDeriv hn w).symm + exact ((hasFTaylorSeriesUpTo_fourierIntegral' L hf h'f).eq_iteratedFDeriv + (mod_cast hn) w).symm end SecondCountableTopology diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean index 6a56975c150b0..d05eed644c71c 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean @@ -368,7 +368,7 @@ theorem contDiff_rpow_const_of_le {p : ℝ} {n : ℕ} (h : ↑n ≤ p) : · exact contDiff_zero.2 (continuous_id.rpow_const fun x => Or.inr <| by simpa using h) · have h1 : 1 ≤ p := le_trans (by simp) h rw [Nat.cast_succ, ← le_sub_iff_add_le] at h - rw [contDiff_succ_iff_deriv, deriv_rpow_const' h1] + rw [show ((n + 1 : ℕ) : ℕ∞) = n + 1 from rfl, contDiff_succ_iff_deriv, deriv_rpow_const' h1] exact ⟨differentiable_rpow_const h1, contDiff_const.mul (ihn h)⟩ theorem contDiffAt_rpow_const_of_le {x p : ℝ} {n : ℕ} (h : ↑n ≤ p) : diff --git a/Mathlib/Data/ENat/Basic.lean b/Mathlib/Data/ENat/Basic.lean index 2c4dde0e12e6a..663a5824cbe31 100644 --- a/Mathlib/Data/ENat/Basic.lean +++ b/Mathlib/Data/ENat/Basic.lean @@ -269,6 +269,15 @@ theorem nat_induction {P : ℕ∞ → Prop} (a : ℕ∞) (h0 : P 0) (hsuc : ∀ · exact htop A · exact A _ +lemma add_one_pos : 0 < n + 1 := + succ_def n ▸ Order.bot_lt_succ n + +lemma add_lt_add_iff_right {k : ℕ∞} (h : k ≠ ⊤) : n + k < m + k ↔ n < m := + WithTop.add_lt_add_iff_right h + +lemma add_lt_add_iff_left {k : ℕ∞} (h : k ≠ ⊤) : k + n < k + m ↔ n < m := + WithTop.add_lt_add_iff_left h + protected lemma exists_nat_gt {n : ℕ∞} (hn : n ≠ ⊤) : ∃ m : ℕ, n < m := by lift n to ℕ using hn obtain ⟨m, hm⟩ := exists_gt n @@ -286,6 +295,8 @@ protected lemma le_sub_of_add_le_left (ha : a ≠ ⊤) : a + b ≤ c → b ≤ c protected lemma sub_sub_cancel (h : a ≠ ⊤) (h2 : b ≤ a) : a - (a - b) = b := (addLECancellable_of_ne_top <| ne_top_of_le_ne_top h tsub_le_self).tsub_tsub_cancel_of_le h2 +section withTop_enat + lemma add_one_natCast_le_withTop_of_lt {m : ℕ} {n : WithTop ℕ∞} (h : m < n) : (m + 1 : ℕ) ≤ n := by match n with | ⊤ => exact le_top @@ -294,9 +305,10 @@ lemma add_one_natCast_le_withTop_of_lt {m : ℕ} {n : WithTop ℕ∞} (h : m < n @[simp] lemma coe_top_add_one : ((⊤ : ℕ∞) : WithTop ℕ∞) + 1 = (⊤ : ℕ∞) := rfl -@[simp] lemma add_one_eq_coe_top_iff : n + 1 = (⊤ : ℕ∞) ↔ n = (⊤ : ℕ∞) := by +@[simp] lemma add_one_eq_coe_top_iff {n : WithTop ℕ∞} : n + 1 = (⊤ : ℕ∞) ↔ n = (⊤ : ℕ∞) := by match n with | ⊤ => exact Iff.rfl + | (⊤ : ℕ∞) => simp | (n : ℕ) => norm_cast; simp only [coe_ne_top, iff_false, ne_eq] @[simp] lemma natCast_ne_coe_top (n : ℕ) : (n : WithTop ℕ∞) ≠ (⊤ : ℕ∞) := nofun @@ -308,13 +320,12 @@ lemma one_le_iff_ne_zero_withTop {n : WithTop ℕ∞} : 1 ≤ n ↔ n ≠ 0 := ⟨fun h ↦ (zero_lt_one.trans_le h).ne', fun h ↦ add_one_natCast_le_withTop_of_lt (pos_iff_ne_zero.mpr h)⟩ -lemma add_one_pos : 0 < n + 1 := - succ_def n ▸ Order.bot_lt_succ n +lemma natCast_le_of_coe_top_le_withTop {N : WithTop ℕ∞} (hN : (⊤ : ℕ∞) ≤ N) (n : ℕ) : n ≤ N := + le_trans (mod_cast le_top) hN -lemma add_lt_add_iff_right {k : ℕ∞} (h : k ≠ ⊤) : n + k < m + k ↔ n < m := - WithTop.add_lt_add_iff_right h +lemma natCast_lt_of_coe_top_le_withTop {N : WithTop ℕ∞} (hN : (⊤ : ℕ∞) ≤ N) (n : ℕ) : n < N := + lt_of_lt_of_le (mod_cast lt_add_one n) (natCast_le_of_coe_top_le_withTop hN (n + 1)) -lemma add_lt_add_iff_left {k : ℕ∞} (h : k ≠ ⊤) : k + n < k + m ↔ n < m := - WithTop.add_lt_add_iff_left h +end withTop_enat end ENat From fb23240e8b4409f2f23bcd07aa73c7451f216413 Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 20 Nov 2024 13:28:39 +0000 Subject: [PATCH 149/829] feat(CI): add a log of the size of the oleans (#19283) Split off from #16020: it only adds a step that computes the sizes of the oleans. --- .github/build.in.yml | 4 ++++ .github/workflows/bors.yml | 4 ++++ .github/workflows/build.yml | 4 ++++ .github/workflows/build_fork.yml | 4 ++++ 4 files changed, 16 insertions(+) diff --git a/.github/build.in.yml b/.github/build.in.yml index 3afe6166a83f4..cf758b28c47f3 100644 --- a/.github/build.in.yml +++ b/.github/build.in.yml @@ -156,6 +156,10 @@ jobs: run: | bash -o pipefail -c "env LEAN_ABORT_ON_PANIC=1 lake build --wfail -KCI" + - name: print the sizes of the oleans + run: | + du .lake/build/lib/Mathlib || echo "This code should be unreachable" + - name: upload cache # We only upload the cache if the build started (whether succeeding, failing, or cancelled) # but not if any earlier step failed or was cancelled. diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index 598668edd66d3..68c1d03a1f5f2 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -166,6 +166,10 @@ jobs: run: | bash -o pipefail -c "env LEAN_ABORT_ON_PANIC=1 lake build --wfail -KCI" + - name: print the sizes of the oleans + run: | + du .lake/build/lib/Mathlib || echo "This code should be unreachable" + - name: upload cache # We only upload the cache if the build started (whether succeeding, failing, or cancelled) # but not if any earlier step failed or was cancelled. diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 4afc15b2f5939..7edf488cd7923 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -173,6 +173,10 @@ jobs: run: | bash -o pipefail -c "env LEAN_ABORT_ON_PANIC=1 lake build --wfail -KCI" + - name: print the sizes of the oleans + run: | + du .lake/build/lib/Mathlib || echo "This code should be unreachable" + - name: upload cache # We only upload the cache if the build started (whether succeeding, failing, or cancelled) # but not if any earlier step failed or was cancelled. diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index 1964290d30bc7..76448210f9cb0 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -170,6 +170,10 @@ jobs: run: | bash -o pipefail -c "env LEAN_ABORT_ON_PANIC=1 lake build --wfail -KCI" + - name: print the sizes of the oleans + run: | + du .lake/build/lib/Mathlib || echo "This code should be unreachable" + - name: upload cache # We only upload the cache if the build started (whether succeeding, failing, or cancelled) # but not if any earlier step failed or was cancelled. From 09242aa0f47b50626ae7552e5695f06fd4068b7a Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Wed, 20 Nov 2024 14:09:17 +0000 Subject: [PATCH 150/829] feat(KrullDimension): height refactor module docstring (#15524) This concludes a refactoring of and API extension for `height`. Notable changes, spread over 28 prior PRs, include: * Everything put into `namespace Order` * Type of `krullDim` changed from `WithBot (WithTop Nat)` to `WithBot ENat`. * Type of `height` changed from `WithBot (WithTop Nat)` to `ENat`. * Definition of `height` via `LTSeries` rather than krullDim * Plenty of lemmas about `height` * Concrete calculations for `height` and `krullDim` This final change updates the module docstring, and adds two `coheight` lemmas corresponding to existing `height` lemmas that slipped through the cracks of juggling multple PRs. From the Carleson project. --- Mathlib/Order/KrullDimension.lean | 30 +++++++++++++++++++++++++++--- 1 file changed, 27 insertions(+), 3 deletions(-) diff --git a/Mathlib/Order/KrullDimension.lean b/Mathlib/Order/KrullDimension.lean index 249a8d2cc0301..1b435cca65877 100644 --- a/Mathlib/Order/KrullDimension.lean +++ b/Mathlib/Order/KrullDimension.lean @@ -19,16 +19,32 @@ In case that `α` is empty, then its Krull dimension is defined to be negative i length of all series `a₀ < a₁ < ... < aₙ` is unbounded, then its Krull dimension is defined to be positive infinity. -For `a : α`, its height (in `ℕ∞`) is defined to be `sup {n | a₀ < a₁ < ... < aₙ ≤ a}` while its +For `a : α`, its height (in `ℕ∞`) is defined to be `sup {n | a₀ < a₁ < ... < aₙ ≤ a}`, while its coheight is defined to be `sup {n | a ≤ a₀ < a₁ < ... < aₙ}` . ## Main results * The Krull dimension is the same as that of the dual order (`krullDim_orderDual`). -* The Krull dimension is the supremum of the heights of the elements (`krullDim_eq_iSup_height`). +* The Krull dimension is the supremum of the heights of the elements (`krullDim_eq_iSup_height`), + or their coheights (`krullDim_eq_iSup_coheight`), or their sums of height and coheight + (`krullDim_eq_iSup_height_add_coheight_of_nonempty`) -* The height is monotone. +* The height in the dual order equals the coheight, and vice versa. + +* The height is monotone (`height_mono`), and strictly monotone if finite (`height_strictMono`). + +* The coheight is antitone (`coheight_anti`), and strictly antitone if finite + (`coheight_strictAnti`). + +* The height is the supremum of the successor of the height of all smaller elements + (`height_eq_iSup_lt_height`). + +* The elements of height zero are the minimal elements (`height_eq_zero`), and the elements of + height `n` are minimal among those of height `≥ n` (`height_eq_coe_iff_minimal_le_height`). + +* Concrete calculations for the height, coheight and Krull dimension in `ℕ`, `ℤ`, `WithTop`, + `WithBot` and `ℕ∞`. ## Design notes @@ -240,6 +256,9 @@ lemma height_mono : Monotone (α := α) height := lemma coheight_anti : Antitone (α := α) coheight := (height_mono (α := αᵒᵈ)).dual_left +@[gcongr] protected lemma _root_.GCongr.coheight_le_coheight (a b : α) (hba : b ≤ a) : + coheight a ≤ coheight b := coheight_anti hba + private lemma height_add_const (a : α) (n : ℕ∞) : height a + n = ⨆ (p : LTSeries α) (_ : p.last = a), p.length + n := by have hne : Nonempty { p : LTSeries α // p.last = a } := ⟨RelSeries.singleton _ a, rfl⟩ @@ -253,6 +272,11 @@ private lemma height_add_const (a : α) (n : ℕ∞) : have := length_le_height_last (p := p.snoc y (by simp [*])) simpa using this +/- For elements of finite height, `coheight` is strictly antitone. -/ +@[gcongr] lemma coheight_strictAnti {x y : α} (hyx : y < x) (hfin : coheight x < ⊤) : + coheight x < coheight y := + height_strictMono (α := αᵒᵈ) hyx hfin + lemma height_le_height_apply_of_strictMono (f : α → β) (hf : StrictMono f) (x : α) : height x ≤ height (f x) := by simp only [height_eq_iSup_last_eq] From edcdc3e1603a3d61b9e9e2cea650beaa663f26bb Mon Sep 17 00:00:00 2001 From: grhkm21 Date: Wed, 20 Nov 2024 14:09:19 +0000 Subject: [PATCH 151/829] feat(UpperHalfPlane): weaken assumptions of multiple atImInfty lemmas (#16015) In this PR, I weaken some assumptions so that they can be applied more widely (more specifically, to my use case :) ) I also renamed some lemmas because they sounded weird. --- Mathlib/Analysis/Asymptotics/Asymptotics.lean | 2 +- .../FunctionsBoundedAtInfty.lean | 23 +++++++++++-------- Mathlib/NumberTheory/ModularForms/Basic.lean | 5 +--- .../EisensteinSeries/IsBoundedAtImInfty.lean | 2 +- .../Order/Filter/ZeroAndBoundedAtFilter.lean | 9 ++++---- 5 files changed, 21 insertions(+), 20 deletions(-) diff --git a/Mathlib/Analysis/Asymptotics/Asymptotics.lean b/Mathlib/Analysis/Asymptotics/Asymptotics.lean index 87c9e471f1eac..c9ef3caef6893 100644 --- a/Mathlib/Analysis/Asymptotics/Asymptotics.lean +++ b/Mathlib/Analysis/Asymptotics/Asymptotics.lean @@ -1149,7 +1149,7 @@ theorem isLittleO_const_iff_isLittleO_one {c : F''} (hc : c ≠ 0) : fun h => h.trans_isBigO <| isBigO_const_const _ hc _⟩ @[simp] -theorem isLittleO_one_iff : f' =o[l] (fun _x => 1 : α → F) ↔ Tendsto f' l (𝓝 0) := by +theorem isLittleO_one_iff {f : α → E'''} : f =o[l] (fun _x => 1 : α → F) ↔ Tendsto f l (𝓝 0) := by simp only [isLittleO_iff, norm_one, mul_one, Metric.nhds_basis_closedBall.tendsto_right_iff, Metric.mem_closedBall, dist_zero_right] diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/FunctionsBoundedAtInfty.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/FunctionsBoundedAtInfty.lean index 7ecb4b7001f40..b784bb4935538 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/FunctionsBoundedAtInfty.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/FunctionsBoundedAtInfty.lean @@ -53,19 +53,22 @@ def zeroAtImInftySubmodule (α : Type*) [NormedField α] : Submodule α (ℍ → def boundedAtImInftySubalgebra (α : Type*) [NormedField α] : Subalgebra α (ℍ → α) := boundedFilterSubalgebra _ atImInfty -nonrec theorem IsBoundedAtImInfty.mul {f g : ℍ → ℂ} (hf : IsBoundedAtImInfty f) - (hg : IsBoundedAtImInfty g) : IsBoundedAtImInfty (f * g) := by - simpa only [Pi.one_apply, mul_one, norm_eq_abs] using hf.mul hg - -theorem bounded_mem (f : ℍ → ℂ) : - IsBoundedAtImInfty f ↔ ∃ M A : ℝ, ∀ z : ℍ, A ≤ im z → abs (f z) ≤ M := by +theorem isBoundedAtImInfty_iff {α : Type*} [Norm α] {f : ℍ → α} : + IsBoundedAtImInfty f ↔ ∃ M A : ℝ, ∀ z : ℍ, A ≤ im z → ‖f z‖ ≤ M := by simp [IsBoundedAtImInfty, BoundedAtFilter, Asymptotics.isBigO_iff, Filter.Eventually, atImInfty_mem] -theorem zero_at_im_infty (f : ℍ → ℂ) : - IsZeroAtImInfty f ↔ ∀ ε : ℝ, 0 < ε → ∃ A : ℝ, ∀ z : ℍ, A ≤ im z → abs (f z) ≤ ε := - (atImInfty_basis.tendsto_iff Metric.nhds_basis_closedBall).trans <| by - simp only [true_and, mem_closedBall_zero_iff]; rfl +@[deprecated (since := "2024-08-27")] alias _root_.bounded_mem := isBoundedAtImInfty_iff + +theorem isZeroAtImInfty_iff {α : Type*} [SeminormedAddGroup α] {f : ℍ → α} : + IsZeroAtImInfty f ↔ ∀ ε : ℝ, 0 < ε → ∃ A : ℝ, ∀ z : ℍ, A ≤ im z → ‖f z‖ ≤ ε := + (atImInfty_basis.tendsto_iff Metric.nhds_basis_closedBall).trans <| by simp + +@[deprecated (since := "2024-08-27")] alias _root_.zero_at_im_infty := isZeroAtImInfty_iff + +theorem IsZeroAtImInfty.isBoundedAtImInfty {α : Type*} [SeminormedAddGroup α] {f : ℍ → α} + (hf : IsZeroAtImInfty f) : IsBoundedAtImInfty f := + hf.boundedAtFilter lemma tendsto_comap_im_ofComplex : Tendsto ofComplex (comap Complex.im atTop) atImInfty := by diff --git a/Mathlib/NumberTheory/ModularForms/Basic.lean b/Mathlib/NumberTheory/ModularForms/Basic.lean index d82d0134a4d49..2a7ff22c6f898 100644 --- a/Mathlib/NumberTheory/ModularForms/Basic.lean +++ b/Mathlib/NumberTheory/ModularForms/Basic.lean @@ -227,10 +227,7 @@ def mul {k_1 k_2 : ℤ} {Γ : Subgroup SL(2, ℤ)} (f : ModularForm Γ k_1) (g : toSlashInvariantForm := f.1.mul g.1 holo' := f.holo'.mul g.holo' bdd_at_infty' A := by - -- Porting note: was `by simpa using ...` - -- `mul_slash_SL2` is no longer a `simp` and `simpa only [mul_slash_SL2] using ...` failed - rw [SlashInvariantForm.coe_mul, mul_slash_SL2] - exact (f.bdd_at_infty' A).mul (g.bdd_at_infty' A) + simpa only [coe_mul, mul_slash_SL2] using (f.bdd_at_infty' A).mul (g.bdd_at_infty' A) @[simp] theorem mul_coe {k_1 k_2 : ℤ} {Γ : Subgroup SL(2, ℤ)} (f : ModularForm Γ k_1) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/IsBoundedAtImInfty.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/IsBoundedAtImInfty.lean index d5fe5f9673fc0..eb74346576451 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/IsBoundedAtImInfty.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/IsBoundedAtImInfty.lean @@ -52,7 +52,7 @@ lemma abs_le_tsum_abs (N : ℕ) (a : Fin 2 → ZMod N) (k : ℤ) (hk : 3 ≤ k) /-- Eisenstein series are bounded at infinity. -/ theorem isBoundedAtImInfty_eisensteinSeries_SIF {N : ℕ+} (a : Fin 2 → ZMod N) {k : ℤ} (hk : 3 ≤ k) (A : SL(2, ℤ)) : IsBoundedAtImInfty ((eisensteinSeries_SIF a k).toFun ∣[k] A) := by - simp_rw [UpperHalfPlane.bounded_mem, eisensteinSeries_SIF] at * + simp_rw [UpperHalfPlane.isBoundedAtImInfty_iff, eisensteinSeries_SIF] at * refine ⟨∑'(x : Fin 2 → ℤ), r ⟨⟨N, 2⟩, Nat.ofNat_pos⟩ ^ (-k) * ‖x‖ ^ (-k), 2, ?_⟩ intro z hz obtain ⟨n, hn⟩ := (ModularGroup_T_zpow_mem_verticalStrip z N.2) diff --git a/Mathlib/Order/Filter/ZeroAndBoundedAtFilter.lean b/Mathlib/Order/Filter/ZeroAndBoundedAtFilter.lean index c7c14ca53e212..fd33cad02fcd3 100644 --- a/Mathlib/Order/Filter/ZeroAndBoundedAtFilter.lean +++ b/Mathlib/Order/Filter/ZeroAndBoundedAtFilter.lean @@ -71,15 +71,16 @@ if `f =O[l] 1`. -/ def BoundedAtFilter [Norm β] (l : Filter α) (f : α → β) : Prop := Asymptotics.IsBigO l f (1 : α → ℝ) -theorem ZeroAtFilter.boundedAtFilter [NormedAddCommGroup β] {l : Filter α} {f : α → β} - (hf : ZeroAtFilter l f) : BoundedAtFilter l f := by - rw [ZeroAtFilter, ← Asymptotics.isLittleO_const_iff (one_ne_zero' ℝ)] at hf - exact hf.isBigO +theorem ZeroAtFilter.boundedAtFilter [SeminormedAddGroup β] {l : Filter α} {f : α → β} + (hf : ZeroAtFilter l f) : BoundedAtFilter l f := + ((Asymptotics.isLittleO_one_iff _).mpr hf).isBigO theorem const_boundedAtFilter [Norm β] (l : Filter α) (c : β) : BoundedAtFilter l (Function.const α c : α → β) := Asymptotics.isBigO_const_const c one_ne_zero l +-- TODO(https://github.com/leanprover-community/mathlib4/issues/19288): Remove all Comm in the next +-- three lemmas. This would require modifying the corresponding general asymptotics lemma. nonrec theorem BoundedAtFilter.add [SeminormedAddCommGroup β] {l : Filter α} {f g : α → β} (hf : BoundedAtFilter l f) (hg : BoundedAtFilter l g) : BoundedAtFilter l (f + g) := by simpa using hf.add hg From f02eb6153c95cc977fdc13456fa9da71348e936b Mon Sep 17 00:00:00 2001 From: FR Date: Wed, 20 Nov 2024 14:09:20 +0000 Subject: [PATCH 152/829] chore: deprecate `Quotient.out'` (#17941) --- Mathlib/Data/Multiset/Basic.lean | 2 +- Mathlib/Data/Quot.lean | 10 ++-- Mathlib/Data/Setoid/Basic.lean | 4 +- Mathlib/Data/Setoid/Partition.lean | 10 ++-- Mathlib/GroupTheory/Complement.lean | 50 ++++++++--------- Mathlib/GroupTheory/Coset/Basic.lean | 17 +++--- Mathlib/GroupTheory/Coset/Defs.lean | 15 +++--- Mathlib/GroupTheory/DoubleCoset.lean | 20 ++++--- Mathlib/GroupTheory/GroupAction/Defs.lean | 8 +-- Mathlib/GroupTheory/GroupAction/Quotient.lean | 53 +++++++++++-------- Mathlib/GroupTheory/Transfer.lean | 14 ++--- .../LinearAlgebra/Projectivization/Basic.lean | 4 +- Mathlib/MeasureTheory/Function/AEEqFun.lean | 2 +- .../MeasureTheory/Measure/Haar/Quotient.lean | 2 +- .../NumberField/Units/DirichletTheorem.lean | 2 +- Mathlib/Order/Antisymmetrization.lean | 2 +- Mathlib/Order/RelIso/Basic.lean | 6 ++- .../HomogeneousLocalization.lean | 12 ++--- Mathlib/RingTheory/Ideal/Quotient/Basic.lean | 4 +- Mathlib/RingTheory/Perfection.lean | 4 +- Mathlib/Topology/Algebra/ClosedSubgroup.lean | 4 +- Mathlib/Topology/UniformSpace/Separation.lean | 2 +- 22 files changed, 136 insertions(+), 111 deletions(-) diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index cebed32a1d12a..a594866b4536d 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -412,7 +412,7 @@ section ToList /-- Produces a list of the elements in the multiset using choice. -/ noncomputable def toList (s : Multiset α) := - s.out' + s.out @[simp, norm_cast] theorem coe_toList (s : Multiset α) : (s.toList : Multiset α) = s := diff --git a/Mathlib/Data/Quot.lean b/Mathlib/Data/Quot.lean index 253f9ac3dded9..9508faa6a1ad6 100644 --- a/Mathlib/Data/Quot.lean +++ b/Mathlib/Data/Quot.lean @@ -727,16 +727,12 @@ protected theorem eq' {s₁ : Setoid α} {a b : α} : protected theorem eq'' {a b : α} : @Quotient.mk'' α s₁ a = Quotient.mk'' b ↔ s₁ a b := Quotient.eq -/-- A version of `Quotient.out` taking `{s₁ : Setoid α}` as an implicit argument instead of an -instance argument. -/ -noncomputable def out' (a : Quotient s₁) : α := - Quotient.out a +@[deprecated (since := "2024-10-19")] alias out' := out -@[simp] -theorem out_eq' (q : Quotient s₁) : Quotient.mk'' q.out' = q := +theorem out_eq' (q : Quotient s₁) : Quotient.mk'' q.out = q := q.out_eq -theorem mk_out' (a : α) : s₁ (Quotient.mk'' a : Quotient s₁).out' a := +theorem mk_out' (a : α) : s₁ (Quotient.mk'' a : Quotient s₁).out a := Quotient.exact (Quotient.out_eq _) section diff --git a/Mathlib/Data/Setoid/Basic.lean b/Mathlib/Data/Setoid/Basic.lean index 5774b37a49f86..2e75e96025a44 100644 --- a/Mathlib/Data/Setoid/Basic.lean +++ b/Mathlib/Data/Setoid/Basic.lean @@ -101,6 +101,8 @@ theorem ker_mk_eq (r : Setoid α) : ker (@Quotient.mk'' _ r) = r := theorem ker_apply_mk_out {f : α → β} (a : α) : f (⟦a⟧ : Quotient (Setoid.ker f)).out = f a := @Quotient.mk_out _ (Setoid.ker f) a +set_option linter.deprecated false in +@[deprecated ker_apply_mk_out (since := "2024-10-19")] theorem ker_apply_mk_out' {f : α → β} (a : α) : f (Quotient.mk _ a : Quotient <| Setoid.ker f).out' = f a := @Quotient.mk_out' _ (Setoid.ker f) a @@ -146,7 +148,7 @@ equivalence relations. -/ @[simps] noncomputable def piQuotientEquiv {ι : Sort*} {α : ι → Sort*} (r : ∀ i, Setoid (α i)) : (∀ i, Quotient (r i)) ≃ Quotient (@piSetoid _ _ r) where - toFun := fun x ↦ Quotient.mk'' fun i ↦ (x i).out' + toFun := fun x ↦ Quotient.mk'' fun i ↦ (x i).out invFun := fun q ↦ Quotient.liftOn' q (fun x i ↦ Quotient.mk'' (x i)) fun x y hxy ↦ by ext i simpa using hxy i diff --git a/Mathlib/Data/Setoid/Partition.lean b/Mathlib/Data/Setoid/Partition.lean index 92f2964a06ba7..1aa847fee7ff8 100644 --- a/Mathlib/Data/Setoid/Partition.lean +++ b/Mathlib/Data/Setoid/Partition.lean @@ -405,7 +405,7 @@ theorem equivQuotient_index : hs.equivQuotient ∘ hs.index = hs.proj := funext hs.equivQuotient_index_apply /-- A map choosing a representative for each element of the quotient associated to an indexed -partition. This is a computable version of `Quotient.out'` using `IndexedPartition.some`. -/ +partition. This is a computable version of `Quotient.out` using `IndexedPartition.some`. -/ def out : hs.Quotient ↪ α := hs.equivQuotient.symm.toEmbedding.trans ⟨hs.some, Function.LeftInverse.injective hs.index_some⟩ @@ -414,9 +414,11 @@ def out : hs.Quotient ↪ α := theorem out_proj (x : α) : hs.out (hs.proj x) = hs.some (hs.index x) := rfl -/-- The indices of `Quotient.out'` and `IndexedPartition.out` are equal. -/ -theorem index_out' (x : hs.Quotient) : hs.index x.out' = hs.index (hs.out x) := - Quotient.inductionOn' x fun x => (Setoid.ker_apply_mk_out' x).trans (hs.index_some _).symm +/-- The indices of `Quotient.out` and `IndexedPartition.out` are equal. -/ +theorem index_out (x : hs.Quotient) : hs.index x.out = hs.index (hs.out x) := + Quotient.inductionOn' x fun x => (Setoid.ker_apply_mk_out x).trans (hs.index_some _).symm + +@[deprecated (since := "2024-10-19")] alias index_out' := index_out /-- This lemma is analogous to `Quotient.out_eq'`. -/ @[simp] diff --git a/Mathlib/GroupTheory/Complement.lean b/Mathlib/GroupTheory/Complement.lean index 5884f8c7afa96..dc8d224d031ad 100644 --- a/Mathlib/GroupTheory/Complement.lean +++ b/Mathlib/GroupTheory/Complement.lean @@ -257,11 +257,11 @@ lemma exists_left_transversal (H : Subgroup G) (g : G) : ∃ S ∈ leftTransversals (H : Set G), g ∈ S := by classical refine - ⟨Set.range (Function.update Quotient.out' _ g), range_mem_leftTransversals fun q => ?_, - Quotient.mk'' g, Function.update_same (Quotient.mk'' g) g Quotient.out'⟩ + ⟨Set.range (Function.update Quotient.out _ g), range_mem_leftTransversals fun q => ?_, + Quotient.mk'' g, Function.update_same (Quotient.mk'' g) g Quotient.out⟩ by_cases hq : q = Quotient.mk'' g - · exact hq.symm ▸ congr_arg _ (Function.update_same (Quotient.mk'' g) g Quotient.out') - · refine (Function.update_noteq ?_ g Quotient.out') ▸ q.out_eq' + · exact hq.symm ▸ congr_arg _ (Function.update_same (Quotient.mk'' g) g Quotient.out) + · refine (Function.update_noteq ?_ g Quotient.out) ▸ q.out_eq' exact hq @[to_additive] @@ -269,11 +269,11 @@ lemma exists_right_transversal (H : Subgroup G) (g : G) : ∃ S ∈ rightTransversals (H : Set G), g ∈ S := by classical refine - ⟨Set.range (Function.update Quotient.out' _ g), range_mem_rightTransversals fun q => ?_, - Quotient.mk'' g, Function.update_same (Quotient.mk'' g) g Quotient.out'⟩ + ⟨Set.range (Function.update Quotient.out _ g), range_mem_rightTransversals fun q => ?_, + Quotient.mk'' g, Function.update_same (Quotient.mk'' g) g Quotient.out⟩ by_cases hq : q = Quotient.mk'' g - · exact hq.symm ▸ congr_arg _ (Function.update_same (Quotient.mk'' g) g Quotient.out') - · exact Eq.trans (congr_arg _ (Function.update_noteq hq g Quotient.out')) q.out_eq' + · exact hq.symm ▸ congr_arg _ (Function.update_same (Quotient.mk'' g) g Quotient.out) + · exact Eq.trans (congr_arg _ (Function.update_noteq hq g Quotient.out)) q.out_eq' /-- Given two subgroups `H' ⊆ H`, there exists a left transversal to `H'` inside `H`. -/ @[to_additive "Given two subgroups `H' ⊆ H`, there exists a transversal to `H'` inside `H`"] @@ -587,11 +587,11 @@ end Action @[to_additive] instance : Inhabited (leftTransversals (H : Set G)) := - ⟨⟨Set.range Quotient.out', range_mem_leftTransversals Quotient.out_eq'⟩⟩ + ⟨⟨Set.range Quotient.out, range_mem_leftTransversals Quotient.out_eq'⟩⟩ @[to_additive] instance : Inhabited (rightTransversals (H : Set G)) := - ⟨⟨Set.range Quotient.out', range_mem_rightTransversals Quotient.out_eq'⟩⟩ + ⟨⟨Set.range Quotient.out, range_mem_rightTransversals Quotient.out_eq'⟩⟩ theorem IsComplement'.isCompl (h : IsComplement' H K) : IsCompl H K := by refine @@ -666,17 +666,17 @@ variable {G : Type u} [Group G] (H : Subgroup G) (g : G) /-- Partition `G ⧸ H` into orbits of the action of `g : G`. -/ noncomputable def quotientEquivSigmaZMod : - G ⧸ H ≃ Σq : orbitRel.Quotient (zpowers g) (G ⧸ H), ZMod (minimalPeriod (g • ·) q.out') := + G ⧸ H ≃ Σq : orbitRel.Quotient (zpowers g) (G ⧸ H), ZMod (minimalPeriod (g • ·) q.out) := (selfEquivSigmaOrbits (zpowers g) (G ⧸ H)).trans - (sigmaCongrRight fun q => orbitZPowersEquiv g q.out') + (sigmaCongrRight fun q => orbitZPowersEquiv g q.out) theorem quotientEquivSigmaZMod_symm_apply (q : orbitRel.Quotient (zpowers g) (G ⧸ H)) - (k : ZMod (minimalPeriod (g • ·) q.out')) : - (quotientEquivSigmaZMod H g).symm ⟨q, k⟩ = g ^ (cast k : ℤ) • q.out' := + (k : ZMod (minimalPeriod (g • ·) q.out)) : + (quotientEquivSigmaZMod H g).symm ⟨q, k⟩ = g ^ (cast k : ℤ) • q.out := rfl theorem quotientEquivSigmaZMod_apply (q : orbitRel.Quotient (zpowers g) (G ⧸ H)) (k : ℤ) : - quotientEquivSigmaZMod H g (g ^ k • q.out') = ⟨q, k⟩ := by + quotientEquivSigmaZMod H g (g ^ k • q.out) = ⟨q, k⟩ := by rw [apply_eq_iff_eq_symm_apply, quotientEquivSigmaZMod_symm_apply, ZMod.coe_intCast, zpow_smul_mod_minimalPeriod] @@ -684,16 +684,16 @@ theorem quotientEquivSigmaZMod_apply (q : orbitRel.Quotient (zpowers g) (G ⧸ H in `G ⧸ H`, an element `g ^ k • q₀` is mapped to `g ^ k • g₀` for a fixed choice of representative `g₀` of `q₀`. -/ noncomputable def transferFunction : G ⧸ H → G := fun q => - g ^ (cast (quotientEquivSigmaZMod H g q).2 : ℤ) * (quotientEquivSigmaZMod H g q).1.out'.out' + g ^ (cast (quotientEquivSigmaZMod H g q).2 : ℤ) * (quotientEquivSigmaZMod H g q).1.out.out theorem transferFunction_apply (q : G ⧸ H) : transferFunction H g q = g ^ (cast (quotientEquivSigmaZMod H g q).2 : ℤ) * - (quotientEquivSigmaZMod H g q).1.out'.out' := + (quotientEquivSigmaZMod H g q).1.out.out := rfl theorem coe_transferFunction (q : G ⧸ H) : ↑(transferFunction H g q) = q := by - rw [transferFunction_apply, ← smul_eq_mul, Quotient.coe_smul_out', + rw [transferFunction_apply, ← smul_eq_mul, Quotient.coe_smul_out, ← quotientEquivSigmaZMod_symm_apply, Sigma.eta, symm_apply_apply] /-- The transfer transversal as a set. Contains elements of the form `g ^ k • g₀` for fixed choices @@ -714,17 +714,17 @@ theorem transferTransversal_apply (q : G ⧸ H) : toEquiv_apply (coe_transferFunction H g) q theorem transferTransversal_apply' (q : orbitRel.Quotient (zpowers g) (G ⧸ H)) - (k : ZMod (minimalPeriod (g • ·) q.out')) : - ↑(toEquiv (transferTransversal H g).2 (g ^ (cast k : ℤ) • q.out')) = - g ^ (cast k : ℤ) * q.out'.out' := by + (k : ZMod (minimalPeriod (g • ·) q.out)) : + ↑(toEquiv (transferTransversal H g).2 (g ^ (cast k : ℤ) • q.out)) = + g ^ (cast k : ℤ) * q.out.out := by rw [transferTransversal_apply, transferFunction_apply, ← quotientEquivSigmaZMod_symm_apply, apply_symm_apply] theorem transferTransversal_apply'' (q : orbitRel.Quotient (zpowers g) (G ⧸ H)) - (k : ZMod (minimalPeriod (g • ·) q.out')) : - ↑(toEquiv (g • transferTransversal H g).2 (g ^ (cast k : ℤ) • q.out')) = - if k = 0 then g ^ minimalPeriod (g • ·) q.out' * q.out'.out' - else g ^ (cast k : ℤ) * q.out'.out' := by + (k : ZMod (minimalPeriod (g • ·) q.out)) : + ↑(toEquiv (g • transferTransversal H g).2 (g ^ (cast k : ℤ) • q.out)) = + if k = 0 then g ^ minimalPeriod (g • ·) q.out * q.out.out + else g ^ (cast k : ℤ) * q.out.out := by rw [smul_apply_eq_smul_apply_inv_smul, transferTransversal_apply, transferFunction_apply, ← mul_smul, ← zpow_neg_one, ← zpow_add, quotientEquivSigmaZMod_apply, smul_eq_mul, ← mul_assoc, ← zpow_one_add, Int.cast_add, Int.cast_neg, Int.cast_one, intCast_cast, cast_id', id, ← diff --git a/Mathlib/GroupTheory/Coset/Basic.lean b/Mathlib/GroupTheory/Coset/Basic.lean index 5caed356835d7..a93c6f92a2109 100644 --- a/Mathlib/GroupTheory/Coset/Basic.lean +++ b/Mathlib/GroupTheory/Coset/Basic.lean @@ -314,10 +314,13 @@ lemma orbit_mk_eq_smul (x : α) : MulAction.orbitRel.Quotient.orbit (x : α ⧸ simpa [mem_smul_set_iff_inv_smul_mem, ← leftRel_apply, Quotient.eq''] using Setoid.comm' _ @[to_additive] -lemma orbit_eq_out'_smul (x : α ⧸ s) : MulAction.orbitRel.Quotient.orbit x = x.out' • s := by +lemma orbit_eq_out_smul (x : α ⧸ s) : MulAction.orbitRel.Quotient.orbit x = x.out • s := by induction x using QuotientGroup.induction_on simp only [orbit_mk_eq_smul, ← eq_class_eq_leftCoset, Quotient.out_eq'] +@[to_additive (attr := deprecated (since := "2024-10-19"))] +alias orbit_eq_out'_smul := orbit_eq_out_smul + end QuotientGroup namespace Subgroup @@ -344,7 +347,7 @@ def rightCosetEquivSubgroup (g : α) : (op g • s : Set α) ≃ s := noncomputable def groupEquivQuotientProdSubgroup : α ≃ (α ⧸ s) × s := calc α ≃ ΣL : α ⧸ s, { x : α // (x : α ⧸ s) = L } := (Equiv.sigmaFiberEquiv QuotientGroup.mk).symm - _ ≃ ΣL : α ⧸ s, (Quotient.out' L • s : Set α) := + _ ≃ ΣL : α ⧸ s, (Quotient.out L • s : Set α) := Equiv.sigmaCongrRight fun L => by rw [← eq_class_eq_leftCoset] show @@ -394,7 +397,7 @@ The constructive version is `quotientEquivProdOfLE'`. -/ @[to_additive (attr := simps!) "If `H ≤ K`, then `G/H ≃ G/K × K/H` nonconstructively. The constructive version is `quotientEquivProdOfLE'`."] noncomputable def quotientEquivProdOfLE (h_le : s ≤ t) : α ⧸ s ≃ (α ⧸ t) × t ⧸ s.subgroupOf t := - quotientEquivProdOfLE' h_le Quotient.out' Quotient.out_eq' + quotientEquivProdOfLE' h_le Quotient.out Quotient.out_eq' /-- If `s ≤ t`, then there is an embedding `s ⧸ H.subgroupOf s ↪ t ⧸ H.subgroupOf t`. -/ @[to_additive "If `s ≤ t`, then there is an embedding @@ -543,11 +546,11 @@ variable [Group α] noncomputable def preimageMkEquivSubgroupProdSet (s : Subgroup α) (t : Set (α ⧸ s)) : QuotientGroup.mk ⁻¹' t ≃ s × t where toFun a := - ⟨⟨((Quotient.out' (QuotientGroup.mk a)) : α)⁻¹ * a, + ⟨⟨((Quotient.out (QuotientGroup.mk a)) : α)⁻¹ * a, leftRel_apply.mp (@Quotient.exact' _ (leftRel s) _ _ <| Quotient.out_eq' _)⟩, ⟨QuotientGroup.mk a, a.2⟩⟩ invFun a := - ⟨Quotient.out' a.2.1 * a.1.1, + ⟨Quotient.out a.2.1 * a.1.1, show QuotientGroup.mk _ ∈ t by rw [mk_mul_of_mem _ a.1.2, out_eq'] exact a.2.2⟩ @@ -559,8 +562,8 @@ open MulAction in @[to_additive "An additive group is made up of a disjoint union of cosets of an additive subgroup."] lemma univ_eq_iUnion_smul (H : Subgroup α) : - (Set.univ (α := α)) = ⋃ x : α ⧸ H, x.out' • (H : Set _) := by - simp_rw [univ_eq_iUnion_orbit H.op, orbit_eq_out'_smul] + (Set.univ (α := α)) = ⋃ x : α ⧸ H, x.out • (H : Set _) := by + simp_rw [univ_eq_iUnion_orbit H.op, orbit_eq_out_smul] rfl end QuotientGroup diff --git a/Mathlib/GroupTheory/Coset/Defs.lean b/Mathlib/GroupTheory/Coset/Defs.lean index f70f9cbbba8eb..f43753df9ac06 100644 --- a/Mathlib/GroupTheory/Coset/Defs.lean +++ b/Mathlib/GroupTheory/Coset/Defs.lean @@ -197,17 +197,20 @@ protected theorem eq {a b : α} : (a : α ⧸ s) = b ↔ a⁻¹ * b ∈ s := @[to_additive (attr := deprecated (since := "2024-08-04"))] alias eq' := QuotientGroup.eq @[to_additive] -theorem out_eq' (a : α ⧸ s) : mk a.out' = a := +theorem out_eq' (a : α ⧸ s) : mk a.out = a := Quotient.out_eq' a variable (s) -/- It can be useful to write `obtain ⟨h, H⟩ := mk_out'_eq_mul ...`, and then `rw [H]` or +/- It can be useful to write `obtain ⟨h, H⟩ := mk_out_eq_mul ...`, and then `rw [H]` or `simp_rw [H]` or `simp only [H]`. In order for `simp_rw` and `simp only` to work, this lemma is - stated in terms of an arbitrary `h : s`, rather than the specific `h = g⁻¹ * (mk g).out'`. -/ -@[to_additive QuotientAddGroup.mk_out'_eq_mul] -theorem mk_out'_eq_mul (g : α) : ∃ h : s, (mk g : α ⧸ s).out' = g * h := - ⟨⟨g⁻¹ * (mk g).out', QuotientGroup.eq.mp (mk g).out_eq'.symm⟩, by rw [mul_inv_cancel_left]⟩ + stated in terms of an arbitrary `h : s`, rather than the specific `h = g⁻¹ * (mk g).out`. -/ +@[to_additive QuotientAddGroup.mk_out_eq_mul] +theorem mk_out_eq_mul (g : α) : ∃ h : s, (mk g : α ⧸ s).out = g * h := + ⟨⟨g⁻¹ * (mk g).out, QuotientGroup.eq.mp (mk g).out_eq'.symm⟩, by rw [mul_inv_cancel_left]⟩ + +@[to_additive (attr := deprecated (since := "2024-10-19")) QuotientAddGroup.mk_out'_eq_mul] +alias mk_out'_eq_mul := mk_out_eq_mul variable {s} {a b : α} diff --git a/Mathlib/GroupTheory/DoubleCoset.lean b/Mathlib/GroupTheory/DoubleCoset.lean index d80688cff9e27..5efccbbe6d6e7 100644 --- a/Mathlib/GroupTheory/DoubleCoset.lean +++ b/Mathlib/GroupTheory/DoubleCoset.lean @@ -101,7 +101,7 @@ theorem rel_bot_eq_right_group_rel (H : Subgroup G) : /-- Create a doset out of an element of `H \ G / K`-/ def quotToDoset (H K : Subgroup G) (q : Quotient (H : Set G) K) : Set G := - doset q.out' H K + doset q.out H K /-- Map from `G` to `H \ G / K`-/ abbrev mk (H K : Subgroup G) (a : G) : Quotient (H : Set G) K := @@ -115,34 +115,38 @@ theorem eq (H K : Subgroup G) (a b : G) : rw [Quotient.eq''] apply rel_iff -theorem out_eq' (H K : Subgroup G) (q : Quotient ↑H ↑K) : mk H K q.out' = q := +theorem out_eq' (H K : Subgroup G) (q : Quotient ↑H ↑K) : mk H K q.out = q := Quotient.out_eq' q -theorem mk_out'_eq_mul (H K : Subgroup G) (g : G) : - ∃ h k : G, h ∈ H ∧ k ∈ K ∧ (mk H K g : Quotient ↑H ↑K).out' = h * g * k := by - have := eq H K (mk H K g : Quotient ↑H ↑K).out' g +theorem mk_out_eq_mul (H K : Subgroup G) (g : G) : + ∃ h k : G, h ∈ H ∧ k ∈ K ∧ (mk H K g : Quotient ↑H ↑K).out = h * g * k := by + have := eq H K (mk H K g : Quotient ↑H ↑K).out g rw [out_eq'] at this obtain ⟨h, h_h, k, hk, T⟩ := this.1 rfl refine ⟨h⁻¹, k⁻¹, H.inv_mem h_h, K.inv_mem hk, eq_mul_inv_of_mul_eq (eq_inv_mul_of_mul_eq ?_)⟩ rw [← mul_assoc, ← T] +@[deprecated (since := "2024-10-19")] alias mk_out'_eq_mul := mk_out_eq_mul + theorem mk_eq_of_doset_eq {H K : Subgroup G} {a b : G} (h : doset a H K = doset b H K) : mk H K a = mk H K b := by rw [eq] exact mem_doset.mp (h.symm ▸ mem_doset_self H K b) -theorem disjoint_out' {H K : Subgroup G} {a b : Quotient H K} : - a ≠ b → Disjoint (doset a.out' H K) (doset b.out' (H : Set G) K) := by +theorem disjoint_out {H K : Subgroup G} {a b : Quotient H K} : + a ≠ b → Disjoint (doset a.out H K) (doset b.out (H : Set G) K) := by contrapose! intro h simpa [out_eq'] using mk_eq_of_doset_eq (eq_of_not_disjoint h) +@[deprecated (since := "2024-10-19")] alias disjoint_out' := disjoint_out + theorem union_quotToDoset (H K : Subgroup G) : ⋃ q, quotToDoset H K q = Set.univ := by ext x simp only [Set.mem_iUnion, quotToDoset, mem_doset, SetLike.mem_coe, exists_prop, Set.mem_univ, iff_true] use mk H K x - obtain ⟨h, k, h3, h4, h5⟩ := mk_out'_eq_mul H K x + obtain ⟨h, k, h3, h4, h5⟩ := mk_out_eq_mul H K x refine ⟨h⁻¹, H.inv_mem h3, k⁻¹, K.inv_mem h4, ?_⟩ simp only [h5, Subgroup.coe_mk, ← mul_assoc, one_mul, inv_mul_cancel, mul_inv_cancel_right] diff --git a/Mathlib/GroupTheory/GroupAction/Defs.lean b/Mathlib/GroupTheory/GroupAction/Defs.lean index e2621efee5a77..3848050964715 100644 --- a/Mathlib/GroupTheory/GroupAction/Defs.lean +++ b/Mathlib/GroupTheory/GroupAction/Defs.lean @@ -431,7 +431,7 @@ nonrec lemma orbitRel.Quotient.orbit_nonempty (x : orbitRel.Quotient G α) : nonrec lemma orbitRel.Quotient.mapsTo_smul_orbit (g : G) (x : orbitRel.Quotient G α) : Set.MapsTo (g • ·) x.orbit x.orbit := by rw [orbitRel.Quotient.orbit_eq_orbit_out x Quotient.out_eq'] - exact mapsTo_smul_orbit g x.out' + exact mapsTo_smul_orbit g x.out @[to_additive] instance (x : orbitRel.Quotient G α) : MulAction G x.orbit where @@ -485,12 +485,12 @@ local notation "Ω" => orbitRel.Quotient G α /-- Decomposition of a type `X` as a disjoint union of its orbits under a group action. This version is expressed in terms of `MulAction.orbitRel.Quotient.orbit` instead of -`MulAction.orbit`, to avoid mentioning `Quotient.out'`. -/ +`MulAction.orbit`, to avoid mentioning `Quotient.out`. -/ @[to_additive "Decomposition of a type `X` as a disjoint union of its orbits under an additive group action. This version is expressed in terms of `AddAction.orbitRel.Quotient.orbit` instead of - `AddAction.orbit`, to avoid mentioning `Quotient.out'`. "] + `AddAction.orbit`, to avoid mentioning `Quotient.out`. "] def selfEquivSigmaOrbits' : α ≃ Σω : Ω, ω.orbit := letI := orbitRel G α calc @@ -503,7 +503,7 @@ def selfEquivSigmaOrbits' : α ≃ Σω : Ω, ω.orbit := @[to_additive "Decomposition of a type `X` as a disjoint union of its orbits under an additive group action."] -def selfEquivSigmaOrbits : α ≃ Σω : Ω, orbit G ω.out' := +def selfEquivSigmaOrbits : α ≃ Σω : Ω, orbit G ω.out := (selfEquivSigmaOrbits' G α).trans <| Equiv.sigmaCongrRight fun _ => Equiv.Set.ofEq <| orbitRel.Quotient.orbit_eq_orbit_out _ Quotient.out_eq' diff --git a/Mathlib/GroupTheory/GroupAction/Quotient.lean b/Mathlib/GroupTheory/GroupAction/Quotient.lean index af70bfbcd8453..4c15d5da76105 100644 --- a/Mathlib/GroupTheory/GroupAction/Quotient.lean +++ b/Mathlib/GroupTheory/GroupAction/Quotient.lean @@ -90,19 +90,30 @@ theorem Quotient.smul_coe [QuotientAction β H] (b : β) (a : α) : rfl @[to_additive (attr := simp)] -theorem Quotient.mk_smul_out' [QuotientAction β H] (b : β) (q : α ⧸ H) : - QuotientGroup.mk (b • q.out') = b • q := by rw [← Quotient.smul_mk, QuotientGroup.out_eq'] +theorem Quotient.mk_smul_out [QuotientAction β H] (b : β) (q : α ⧸ H) : + QuotientGroup.mk (b • q.out) = b • q := by rw [← Quotient.smul_mk, QuotientGroup.out_eq'] -- Porting note: removed simp attribute, simp can prove this @[to_additive] -theorem Quotient.coe_smul_out' [QuotientAction β H] (b : β) (q : α ⧸ H) : ↑(b • q.out') = b • q := - Quotient.mk_smul_out' H b q +theorem Quotient.coe_smul_out [QuotientAction β H] (b : β) (q : α ⧸ H) : ↑(b • q.out) = b • q := + Quotient.mk_smul_out H b q -theorem _root_.QuotientGroup.out'_conj_pow_minimalPeriod_mem (a : α) (q : α ⧸ H) : - q.out'⁻¹ * a ^ Function.minimalPeriod (a • ·) q * q.out' ∈ H := by - rw [mul_assoc, ← QuotientGroup.eq, QuotientGroup.out_eq', ← smul_eq_mul, Quotient.mk_smul_out', +theorem _root_.QuotientGroup.out_conj_pow_minimalPeriod_mem (a : α) (q : α ⧸ H) : + q.out⁻¹ * a ^ Function.minimalPeriod (a • ·) q * q.out ∈ H := by + rw [mul_assoc, ← QuotientGroup.eq, QuotientGroup.out_eq', ← smul_eq_mul, Quotient.mk_smul_out, eq_comm, pow_smul_eq_iff_minimalPeriod_dvd] +@[to_additive (attr := deprecated (since := "2024-10-19"))] +alias Quotient.mk_smul_out' := Quotient.mk_smul_out + +-- Porting note: removed simp attribute, simp can prove this +@[to_additive (attr := deprecated (since := "2024-10-19"))] +alias Quotient.coe_smul_out' := Quotient.coe_smul_out + +@[deprecated (since := "2024-10-19")] +alias _root_.QuotientGroup.out'_conj_pow_minimalPeriod_mem := + QuotientGroup.out_conj_pow_minimalPeriod_mem + end QuotientAction open QuotientGroup @@ -192,14 +203,14 @@ local notation "Ω" => Quotient <| orbitRel α β /-- **Class formula** : given `G` a group acting on `X` and `φ` a function mapping each orbit of `X` under this action (that is, each element of the quotient of `X` by the relation `orbitRel G X`) to an element in this orbit, this gives a (noncomputable) bijection between `X` and the disjoint union -of `G/Stab(φ(ω))` over all orbits `ω`. In most cases you'll want `φ` to be `Quotient.out'`, so we +of `G/Stab(φ(ω))` over all orbits `ω`. In most cases you'll want `φ` to be `Quotient.out`, so we provide `MulAction.selfEquivSigmaOrbitsQuotientStabilizer'` as a special case. -/ @[to_additive "**Class formula** : given `G` an additive group acting on `X` and `φ` a function mapping each orbit of `X` under this action (that is, each element of the quotient of `X` by the relation `orbit_rel G X`) to an element in this orbit, this gives a (noncomputable) bijection between `X` and the disjoint union of `G/Stab(φ(ω))` over all orbits `ω`. In most - cases you'll want `φ` to be `Quotient.out'`, so we provide + cases you'll want `φ` to be `Quotient.out`, so we provide `AddAction.selfEquivSigmaOrbitsQuotientStabilizer'` as a special case. "] noncomputable def selfEquivSigmaOrbitsQuotientStabilizer' {φ : Ω → β} (hφ : LeftInverse Quotient.mk'' φ) : β ≃ Σω : Ω, α ⧸ stabilizer α (φ ω) := @@ -212,11 +223,11 @@ noncomputable def selfEquivSigmaOrbitsQuotientStabilizer' {φ : Ω → β} /-- **Class formula** for a finite group acting on a finite type. See `MulAction.card_eq_sum_card_group_div_card_stabilizer` for a specialized version using -`Quotient.out'`. -/ +`Quotient.out`. -/ @[to_additive "**Class formula** for a finite group acting on a finite type. See `AddAction.card_eq_sum_card_addGroup_div_card_stabilizer` for a specialized version using - `Quotient.out'`."] + `Quotient.out`."] theorem card_eq_sum_card_group_div_card_stabilizer' [Fintype α] [Fintype β] [Fintype Ω] [∀ b : β, Fintype <| stabilizer α b] {φ : Ω → β} (hφ : LeftInverse Quotient.mk'' φ) : Fintype.card β = ∑ ω : Ω, Fintype.card α / Fintype.card (stabilizer α (φ ω)) := by @@ -231,18 +242,18 @@ theorem card_eq_sum_card_group_div_card_stabilizer' [Fintype α] [Fintype β] [F Fintype.card_congr (selfEquivSigmaOrbitsQuotientStabilizer' α β hφ)] /-- **Class formula**. This is a special case of -`MulAction.self_equiv_sigma_orbits_quotient_stabilizer'` with `φ = Quotient.out'`. -/ +`MulAction.self_equiv_sigma_orbits_quotient_stabilizer'` with `φ = Quotient.out`. -/ @[to_additive "**Class formula**. This is a special case of - `AddAction.self_equiv_sigma_orbits_quotient_stabilizer'` with `φ = Quotient.out'`. "] -noncomputable def selfEquivSigmaOrbitsQuotientStabilizer : β ≃ Σω : Ω, α ⧸ stabilizer α ω.out' := + `AddAction.self_equiv_sigma_orbits_quotient_stabilizer'` with `φ = Quotient.out`. "] +noncomputable def selfEquivSigmaOrbitsQuotientStabilizer : β ≃ Σω : Ω, α ⧸ stabilizer α ω.out := selfEquivSigmaOrbitsQuotientStabilizer' α β Quotient.out_eq' /-- **Class formula** for a finite group acting on a finite type. -/ @[to_additive "**Class formula** for a finite group acting on a finite type."] theorem card_eq_sum_card_group_div_card_stabilizer [Fintype α] [Fintype β] [Fintype Ω] [∀ b : β, Fintype <| stabilizer α b] : - Fintype.card β = ∑ ω : Ω, Fintype.card α / Fintype.card (stabilizer α ω.out') := + Fintype.card β = ∑ ω : Ω, Fintype.card α / Fintype.card (stabilizer α ω.out) := card_eq_sum_card_group_div_card_stabilizer' α β Quotient.out_eq' /-- **Burnside's lemma** : a (noncomputable) bijection between the disjoint union of all @@ -259,16 +270,16 @@ noncomputable def sigmaFixedByEquivOrbitsProdGroup : (Σa : α, fixedBy β a) _ ≃ { ba : β × α // ba.2 • ba.1 = ba.1 } := (Equiv.prodComm α β).subtypeEquiv fun _ => Iff.rfl _ ≃ Σb : β, stabilizer α b := Equiv.subtypeProdEquivSigmaSubtype fun (b : β) a => a ∈ stabilizer α b - _ ≃ Σωb : Σω : Ω, orbit α ω.out', stabilizer α (ωb.2 : β) := + _ ≃ Σωb : Σω : Ω, orbit α ω.out, stabilizer α (ωb.2 : β) := (selfEquivSigmaOrbits α β).sigmaCongrLeft' - _ ≃ Σω : Ω, Σb : orbit α ω.out', stabilizer α (b : β) := - Equiv.sigmaAssoc fun (ω : Ω) (b : orbit α ω.out') => stabilizer α (b : β) - _ ≃ Σω : Ω, Σ _ : orbit α ω.out', stabilizer α ω.out' := + _ ≃ Σω : Ω, Σb : orbit α ω.out, stabilizer α (b : β) := + Equiv.sigmaAssoc fun (ω : Ω) (b : orbit α ω.out) => stabilizer α (b : β) + _ ≃ Σω : Ω, Σ _ : orbit α ω.out, stabilizer α ω.out := Equiv.sigmaCongrRight fun _ => Equiv.sigmaCongrRight fun ⟨_, hb⟩ => (stabilizerEquivStabilizerOfOrbitRel hb).toEquiv - _ ≃ Σω : Ω, orbit α ω.out' × stabilizer α ω.out' := + _ ≃ Σω : Ω, orbit α ω.out × stabilizer α ω.out := Equiv.sigmaCongrRight fun _ => Equiv.sigmaEquivProd _ _ - _ ≃ Σ _ : Ω, α := Equiv.sigmaCongrRight fun ω => orbitProdStabilizerEquivGroup α ω.out' + _ ≃ Σ _ : Ω, α := Equiv.sigmaCongrRight fun ω => orbitProdStabilizerEquivGroup α ω.out _ ≃ Ω × α := Equiv.sigmaEquivProd Ω α /-- **Burnside's lemma** : given a finite group `G` acting on a set `X`, the average number of diff --git a/Mathlib/GroupTheory/Transfer.lean b/Mathlib/GroupTheory/Transfer.lean index ce1d169cd838e..8b40c9d992fa9 100644 --- a/Mathlib/GroupTheory/Transfer.lean +++ b/Mathlib/GroupTheory/Transfer.lean @@ -103,8 +103,8 @@ theorem transfer_eq_prod_quotient_orbitRel_zpowers_quot [FiniteIndex H] (g : G) transfer ϕ g = ∏ q : Quotient (orbitRel (zpowers g) (G ⧸ H)), ϕ - ⟨q.out'.out'⁻¹ * g ^ Function.minimalPeriod (g • ·) q.out' * q.out'.out', - QuotientGroup.out'_conj_pow_minimalPeriod_mem H g q.out'⟩ := by + ⟨q.out.out⁻¹ * g ^ Function.minimalPeriod (g • ·) q.out * q.out.out, + QuotientGroup.out_conj_pow_minimalPeriod_mem H g q.out⟩ := by classical letI := H.fintypeQuotientOfFiniteIndex calc @@ -115,7 +115,7 @@ theorem transfer_eq_prod_quotient_orbitRel_zpowers_quot [FiniteIndex H] (g : G) refine Fintype.prod_congr _ _ (fun q => ?_) simp only [quotientEquivSigmaZMod_symm_apply, transferTransversal_apply', transferTransversal_apply''] - rw [Fintype.prod_eq_single (0 : ZMod (Function.minimalPeriod (g • ·) q.out')) _] + rw [Fintype.prod_eq_single (0 : ZMod (Function.minimalPeriod (g • ·) q.out)) _] · simp only [if_pos, ZMod.cast_zero, zpow_zero, one_mul, mul_assoc] · intro k hk simp only [if_neg hk, inv_mul_cancel] @@ -133,11 +133,11 @@ theorem transfer_eq_pow_aux (g : G) replace key : ∀ (k : ℕ) (g₀ : G), g₀⁻¹ * g ^ k * g₀ ∈ H → g ^ k ∈ H := fun k g₀ hk => (congr_arg (· ∈ H) (key k g₀ hk)).mp hk replace key : ∀ q : G ⧸ H, g ^ Function.minimalPeriod (g • ·) q ∈ H := fun q => - key (Function.minimalPeriod (g • ·) q) q.out' - (QuotientGroup.out'_conj_pow_minimalPeriod_mem H g q) + key (Function.minimalPeriod (g • ·) q) q.out + (QuotientGroup.out_conj_pow_minimalPeriod_mem H g q) let f : Quotient (orbitRel (zpowers g) (G ⧸ H)) → zpowers g := fun q => - (⟨g, mem_zpowers g⟩ : zpowers g) ^ Function.minimalPeriod (g • ·) q.out' - have hf : ∀ q, f q ∈ H.subgroupOf (zpowers g) := fun q => key q.out' + (⟨g, mem_zpowers g⟩ : zpowers g) ^ Function.minimalPeriod (g • ·) q.out + have hf : ∀ q, f q ∈ H.subgroupOf (zpowers g) := fun q => key q.out replace key := Subgroup.prod_mem (H.subgroupOf (zpowers g)) fun q (_ : q ∈ Finset.univ) => hf q simpa only [f, minimalPeriod_eq_card, Finset.prod_pow_eq_pow_sum, Fintype.card_sigma, diff --git a/Mathlib/LinearAlgebra/Projectivization/Basic.lean b/Mathlib/LinearAlgebra/Projectivization/Basic.lean index 9a12ad4ccb657..38496cfef13f4 100644 --- a/Mathlib/LinearAlgebra/Projectivization/Basic.lean +++ b/Mathlib/LinearAlgebra/Projectivization/Basic.lean @@ -84,10 +84,10 @@ protected lemma lift_mk {α : Type*} (f : { v : V // v ≠ 0 } → α) /-- Choose a representative of `v : Projectivization K V` in `V`. -/ protected noncomputable def rep (v : ℙ K V) : V := - v.out' + v.out theorem rep_nonzero (v : ℙ K V) : v.rep ≠ 0 := - v.out'.2 + v.out.2 @[simp] theorem mk_rep (v : ℙ K V) : mk K v.rep v.rep_nonzero = v := Quotient.out_eq' _ diff --git a/Mathlib/MeasureTheory/Function/AEEqFun.lean b/Mathlib/MeasureTheory/Function/AEEqFun.lean index 86720959dec25..38c86bd0a94b3 100644 --- a/Mathlib/MeasureTheory/Function/AEEqFun.lean +++ b/Mathlib/MeasureTheory/Function/AEEqFun.lean @@ -124,7 +124,7 @@ then we choose that one. -/ def cast (f : α →ₘ[μ] β) : α → β := if h : ∃ (b : β), f = mk (const α b) aestronglyMeasurable_const then const α <| Classical.choose h else - AEStronglyMeasurable.mk _ (Quotient.out' f : { f : α → β // AEStronglyMeasurable f μ }).2 + AEStronglyMeasurable.mk _ (Quotient.out f : { f : α → β // AEStronglyMeasurable f μ }).2 /-- A measurable representative of an `AEEqFun` [f] -/ instance instCoeFun : CoeFun (α →ₘ[μ] β) fun _ => α → β := ⟨cast⟩ diff --git a/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean b/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean index e552fa14c3286..a782f7e429aeb 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean @@ -134,7 +134,7 @@ lemma MeasureTheory.QuotientMeasureEqMeasurePreimage.mulInvariantMeasure_quotien obtain ⟨x₁, h⟩ := @Quotient.exists_rep _ (QuotientGroup.leftRel Γ) x convert measure_preimage_smul μ x₁ A using 1 · rw [← h, Measure.map_apply (measurable_const_mul _) hA] - simp [← MulAction.Quotient.coe_smul_out', ← Quotient.mk''_eq_mk] + simp [← MulAction.Quotient.coe_smul_out, ← Quotient.mk''_eq_mk] exact smulInvariantMeasure_quotient ν variable [Countable Γ] [IsMulRightInvariant ν] [SigmaFinite ν] diff --git a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean index c735c10548900..31c16252e0c0d 100644 --- a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean +++ b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean @@ -465,7 +465,7 @@ def basisUnitLattice : Basis (Fin (rank K)) ℤ (unitLattice K) := units in `basisModTorsion`. -/ def fundSystem : Fin (rank K) → (𝓞 K)ˣ := -- `:)` prevents the `⧸` decaying to a quotient by `leftRel` when we unfold this later - fun i => Quotient.out' (Additive.toMul (basisModTorsion K i):) + fun i => Quotient.out (Additive.toMul (basisModTorsion K i):) theorem fundSystem_mk (i : Fin (rank K)) : Additive.ofMul (QuotientGroup.mk (fundSystem K i)) = (basisModTorsion K i) := by diff --git a/Mathlib/Order/Antisymmetrization.lean b/Mathlib/Order/Antisymmetrization.lean index d098ac0bfffb1..75caf3593c355 100644 --- a/Mathlib/Order/Antisymmetrization.lean +++ b/Mathlib/Order/Antisymmetrization.lean @@ -88,7 +88,7 @@ def toAntisymmetrization : α → Antisymmetrization α r := /-- Get a representative from the antisymmetrization. -/ noncomputable def ofAntisymmetrization : Antisymmetrization α r → α := - Quotient.out' + Quotient.out instance [Inhabited α] : Inhabited (Antisymmetrization α r) := by unfold Antisymmetrization; infer_instance diff --git a/Mathlib/Order/RelIso/Basic.lean b/Mathlib/Order/RelIso/Basic.lean index 721ef0208ee6d..cc0083a0b2676 100644 --- a/Mathlib/Order/RelIso/Basic.lean +++ b/Mathlib/Order/RelIso/Basic.lean @@ -363,13 +363,17 @@ noncomputable def Quotient.outRelEmbedding {_ : Setoid α} {r : α → α → Pr refine @fun x y => Quotient.inductionOn₂ x y fun a b => ?_ apply iff_iff_eq.2 (H _ _ _ _ _ _) <;> apply Quotient.mk_out⟩ +set_option linter.deprecated false in /-- `Quotient.out'` as a relation embedding between the lift of a relation and the relation. -/ -@[simps] +@[deprecated Quotient.outRelEmbedding (since := "2024-10-19"), simps] noncomputable def Quotient.out'RelEmbedding {_ : Setoid α} {r : α → α → Prop} (H : ∀ (a₁ b₁ a₂ b₂ : α), a₁ ≈ a₂ → b₁ ≈ b₂ → r a₁ b₁ = r a₂ b₂) : (fun a b => Quotient.liftOn₂' a b r H) ↪r r := { Quotient.outRelEmbedding H with toFun := Quotient.out' } +attribute [deprecated Quotient.outRelEmbedding_apply (since := "2024-10-19")] + Quotient.out'RelEmbedding_apply + @[simp] theorem acc_lift₂_iff {_ : Setoid α} {r : α → α → Prop} {H : ∀ (a₁ b₁ a₂ b₂ : α), a₁ ≈ a₂ → b₁ ≈ b₂ → r a₁ b₁ = r a₂ b₂} {a} : diff --git a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean index 553951dd2ce18..dcab8f6f958a6 100644 --- a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean +++ b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean @@ -483,25 +483,25 @@ variable {𝒜} {x} /-- Numerator of an element in `HomogeneousLocalization x`. -/ def num (f : HomogeneousLocalization 𝒜 x) : A := - (Quotient.out' f).num + (Quotient.out f).num /-- Denominator of an element in `HomogeneousLocalization x`. -/ def den (f : HomogeneousLocalization 𝒜 x) : A := - (Quotient.out' f).den + (Quotient.out f).den /-- For an element in `HomogeneousLocalization x`, degree is the natural number `i` such that `𝒜 i` contains both numerator and denominator. -/ def deg (f : HomogeneousLocalization 𝒜 x) : ι := - (Quotient.out' f).deg + (Quotient.out f).deg theorem den_mem (f : HomogeneousLocalization 𝒜 x) : f.den ∈ x := - (Quotient.out' f).den_mem + (Quotient.out f).den_mem theorem num_mem_deg (f : HomogeneousLocalization 𝒜 x) : f.num ∈ 𝒜 f.deg := - (Quotient.out' f).num.2 + (Quotient.out f).num.2 theorem den_mem_deg (f : HomogeneousLocalization 𝒜 x) : f.den ∈ 𝒜 f.deg := - (Quotient.out' f).den.2 + (Quotient.out f).den.2 theorem eq_num_div_den (f : HomogeneousLocalization 𝒜 x) : f.val = Localization.mk f.num ⟨f.den, f.den_mem⟩ := diff --git a/Mathlib/RingTheory/Ideal/Quotient/Basic.lean b/Mathlib/RingTheory/Ideal/Quotient/Basic.lean index 190587bcc84f3..244da7d3aada9 100644 --- a/Mathlib/RingTheory/Ideal/Quotient/Basic.lean +++ b/Mathlib/RingTheory/Ideal/Quotient/Basic.lean @@ -191,7 +191,7 @@ noncomputable def piQuotEquiv : ((ι → R) ⧸ I.pi ι) ≃ₗ[R ⧸ I] ι → funext fun i => (Submodule.Quotient.eq' _).2 (QuotientAddGroup.leftRel_apply.mp hab i) map_add' := by rintro ⟨_⟩ ⟨_⟩; rfl map_smul' := by rintro ⟨_⟩ ⟨_⟩; rfl - invFun := fun x ↦ Ideal.Quotient.mk (I.pi ι) fun i ↦ Quotient.out' (x i) + invFun := fun x ↦ Ideal.Quotient.mk (I.pi ι) fun i ↦ Quotient.out (x i) left_inv := by rintro ⟨x⟩ exact Ideal.Quotient.eq.2 fun i => Ideal.Quotient.eq.1 (Quotient.out_eq' _) @@ -216,7 +216,7 @@ end Pi open scoped Pointwise in /-- A ring is made up of a disjoint union of cosets of an ideal. -/ lemma univ_eq_iUnion_image_add {R : Type*} [Ring R] (I : Ideal R) : - (Set.univ (α := R)) = ⋃ x : R ⧸ I, x.out' +ᵥ (I : Set R) := + (Set.univ (α := R)) = ⋃ x : R ⧸ I, x.out +ᵥ (I : Set R) := QuotientAddGroup.univ_eq_iUnion_vadd I.toAddSubgroup lemma _root_.Finite.of_finite_quot_finite_ideal {R : Type*} [Ring R] {I : Ideal R} diff --git a/Mathlib/RingTheory/Perfection.lean b/Mathlib/RingTheory/Perfection.lean index 7db1bdda6d237..7d7e45a3a1929 100644 --- a/Mathlib/RingTheory/Perfection.lean +++ b/Mathlib/RingTheory/Perfection.lean @@ -349,13 +349,13 @@ attribute [local instance] Classical.dec /-- For a field `K` with valuation `v : K → ℝ≥0` and ring of integers `O`, a function `O/(p) → ℝ≥0` that sends `0` to `0` and `x + (p)` to `v(x)` as long as `x ∉ (p)`. -/ noncomputable def preVal (x : ModP K v O hv p) : ℝ≥0 := - if x = 0 then 0 else v (algebraMap O K x.out') + if x = 0 then 0 else v (algebraMap O K x.out) variable {K v O hv p} theorem preVal_mk {x : O} (hx : (Ideal.Quotient.mk _ x : ModP K v O hv p) ≠ 0) : preVal K v O hv p (Ideal.Quotient.mk _ x) = v (algebraMap O K x) := by - obtain ⟨r, hr⟩ : ∃ (a : O), a * (p : O) = (Quotient.mk'' x).out' - x := + obtain ⟨r, hr⟩ : ∃ (a : O), a * (p : O) = (Quotient.mk'' x).out - x := Ideal.mem_span_singleton'.1 <| Ideal.Quotient.eq.1 <| Quotient.sound' <| Quotient.mk_out' _ refine (if_neg hx).trans (v.map_eq_of_sub_lt <| lt_of_not_le ?_) erw [← RingHom.map_sub, ← hr, hv.le_iff_dvd] diff --git a/Mathlib/Topology/Algebra/ClosedSubgroup.lean b/Mathlib/Topology/Algebra/ClosedSubgroup.lean index 543b8793ad0fa..f3cf42ab0809f 100644 --- a/Mathlib/Topology/Algebra/ClosedSubgroup.lean +++ b/Mathlib/Topology/Algebra/ClosedSubgroup.lean @@ -102,7 +102,7 @@ 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) + ↦ IsClosed.smul h (Quotient.out x.1) ext x refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ · have : QuotientGroup.mk 1 ≠ QuotientGroup.mk (s := H) x := by @@ -110,7 +110,7 @@ lemma isOpen_of_isClosed_of_finiteIndex (H : Subgroup G) [H.FiniteIndex] 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 + (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⟩ diff --git a/Mathlib/Topology/UniformSpace/Separation.lean b/Mathlib/Topology/UniformSpace/Separation.lean index 331e192d249dc..9b3a324dc67df 100644 --- a/Mathlib/Topology/UniformSpace/Separation.lean +++ b/Mathlib/Topology/UniformSpace/Separation.lean @@ -253,7 +253,7 @@ open Classical in TODO: unify with `SeparationQuotient.lift`. -/ def lift' [T0Space β] (f : α → β) : SeparationQuotient α → β := if hc : UniformContinuous f then lift f fun _ _ h => (h.map hc.continuous).eq - else fun x => f (Nonempty.some ⟨x.out'⟩) + else fun x => f (Nonempty.some ⟨x.out⟩) theorem lift'_mk [T0Space β] {f : α → β} (h : UniformContinuous f) (a : α) : lift' f (mk a) = f a := by rw [lift', dif_pos h, lift_mk] From 185ec9cdb8e7dc5e7ad15854962475f18259c484 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Wed, 20 Nov 2024 14:53:32 +0000 Subject: [PATCH 153/829] feat: the Lie bracket of vector fields in vector spaces (#18852) --- Mathlib.lean | 1 + Mathlib/Analysis/Calculus/VectorField.lean | 343 +++++++++++++++++++++ 2 files changed, 344 insertions(+) create mode 100644 Mathlib/Analysis/Calculus/VectorField.lean diff --git a/Mathlib.lean b/Mathlib.lean index 2c342449a9a42..ec30d813df6af 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1172,6 +1172,7 @@ import Mathlib.Analysis.Calculus.SmoothSeries import Mathlib.Analysis.Calculus.TangentCone import Mathlib.Analysis.Calculus.Taylor import Mathlib.Analysis.Calculus.UniformLimitsDeriv +import Mathlib.Analysis.Calculus.VectorField import Mathlib.Analysis.Complex.AbelLimit import Mathlib.Analysis.Complex.AbsMax import Mathlib.Analysis.Complex.Angle diff --git a/Mathlib/Analysis/Calculus/VectorField.lean b/Mathlib/Analysis/Calculus/VectorField.lean new file mode 100644 index 0000000000000..d4e54bc4485f9 --- /dev/null +++ b/Mathlib/Analysis/Calculus/VectorField.lean @@ -0,0 +1,343 @@ +/- +Copyright (c) 2024 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel +-/ +import Mathlib.Analysis.Calculus.FDeriv.Symmetric + +/-! +# Vector fields in vector spaces + +We study functions of the form `V : E → E` on a vector space, thinking of these as vector fields. +We define several notions in this context, with the aim to generalize them to vector fields on +manifolds. + +Notably, we define the pullback of a vector field under a map, as +`VectorField.pullback 𝕜 f V x := (fderiv 𝕜 f x).inverse (V (f x))` (together with the same notion +within a set). + +In addition to comprehensive API on this notion, the main result is the following: +* `VectorField.leibniz_identity_lieBracket` is the Leibniz + identity `[U, [V, W]] = [[U, V], W] + [V, [U, W]]`. + +-/ + +open Set +open scoped Topology + +noncomputable section + +variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] + {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] + {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] + {G : Type*} [NormedAddCommGroup G] [NormedSpace 𝕜 G] + {V W V₁ W₁ : E → E} {s t : Set E} {x : E} + +/-! +### The Lie bracket of vector fields in a vector space + +We define the Lie bracket of two vector fields, and call it `lieBracket 𝕜 V W x`. We also define +a version localized to sets, `lieBracketWithin 𝕜 V W s x`. We copy the relevant API +of `fderivWithin` and `fderiv` for these notions to get a comprehensive API. +-/ + +namespace VectorField + +variable (𝕜) in +/-- The Lie bracket `[V, W] (x)` of two vector fields at a point, defined as +`DW(x) (V x) - DV(x) (W x)`. -/ +def lieBracket (V W : E → E) (x : E) : E := + fderiv 𝕜 W x (V x) - fderiv 𝕜 V x (W x) + +variable (𝕜) in +/-- The Lie bracket `[V, W] (x)` of two vector fields within a set at a point, defined as +`DW(x) (V x) - DV(x) (W x)` where the derivatives are taken inside `s`. -/ +def lieBracketWithin (V W : E → E) (s : Set E) (x : E) : E := + fderivWithin 𝕜 W s x (V x) - fderivWithin 𝕜 V s x (W x) + +lemma lieBracket_eq : + lieBracket 𝕜 V W = fun x ↦ fderiv 𝕜 W x (V x) - fderiv 𝕜 V x (W x) := rfl + +lemma lieBracketWithin_eq : + lieBracketWithin 𝕜 V W s = + fun x ↦ fderivWithin 𝕜 W s x (V x) - fderivWithin 𝕜 V s x (W x) := rfl + +@[simp] +theorem lieBracketWithin_univ : lieBracketWithin 𝕜 V W univ = lieBracket 𝕜 V W := by + ext1 x + simp [lieBracketWithin, lieBracket] + +lemma lieBracketWithin_eq_zero_of_eq_zero (hV : V x = 0) (hW : W x = 0) : + lieBracketWithin 𝕜 V W s x = 0 := by + simp [lieBracketWithin, hV, hW] + +lemma lieBracket_eq_zero_of_eq_zero (hV : V x = 0) (hW : W x = 0) : + lieBracket 𝕜 V W x = 0 := by + simp [lieBracket, hV, hW] + +lemma lieBracketWithin_smul_left {c : 𝕜} (hV : DifferentiableWithinAt 𝕜 V s x) + (hs : UniqueDiffWithinAt 𝕜 s x) : + lieBracketWithin 𝕜 (c • V) W s x = + c • lieBracketWithin 𝕜 V W s x := by + simp only [lieBracketWithin, Pi.add_apply, map_add, Pi.smul_apply, map_smul, smul_sub] + rw [fderivWithin_const_smul' hs hV] + rfl + +lemma lieBracket_smul_left {c : 𝕜} (hV : DifferentiableAt 𝕜 V x) : + lieBracket 𝕜 (c • V) W x = c • lieBracket 𝕜 V W x := by + simp only [← differentiableWithinAt_univ, ← lieBracketWithin_univ] at hV ⊢ + exact lieBracketWithin_smul_left hV uniqueDiffWithinAt_univ + +lemma lieBracketWithin_smul_right {c : 𝕜} (hW : DifferentiableWithinAt 𝕜 W s x) + (hs : UniqueDiffWithinAt 𝕜 s x) : + lieBracketWithin 𝕜 V (c • W) s x = + c • lieBracketWithin 𝕜 V W s x := by + simp only [lieBracketWithin, Pi.add_apply, map_add, Pi.smul_apply, map_smul, smul_sub] + rw [fderivWithin_const_smul' hs hW] + rfl + +lemma lieBracket_smul_right {c : 𝕜} (hW : DifferentiableAt 𝕜 W x) : + lieBracket 𝕜 V (c • W) x = c • lieBracket 𝕜 V W x := by + simp only [← differentiableWithinAt_univ, ← lieBracketWithin_univ] at hW ⊢ + exact lieBracketWithin_smul_right hW uniqueDiffWithinAt_univ + +lemma lieBracketWithin_add_left (hV : DifferentiableWithinAt 𝕜 V s x) + (hV₁ : DifferentiableWithinAt 𝕜 V₁ s x) (hs : UniqueDiffWithinAt 𝕜 s x) : + lieBracketWithin 𝕜 (V + V₁) W s x = + lieBracketWithin 𝕜 V W s x + lieBracketWithin 𝕜 V₁ W s x := by + simp only [lieBracketWithin, Pi.add_apply, map_add] + rw [fderivWithin_add' hs hV hV₁, ContinuousLinearMap.add_apply] + abel + +lemma lieBracket_add_left (hV : DifferentiableAt 𝕜 V x) (hV₁ : DifferentiableAt 𝕜 V₁ x) : + lieBracket 𝕜 (V + V₁) W x = + lieBracket 𝕜 V W x + lieBracket 𝕜 V₁ W x := by + simp only [lieBracket, Pi.add_apply, map_add] + rw [fderiv_add' hV hV₁, ContinuousLinearMap.add_apply] + abel + +lemma lieBracketWithin_add_right (hW : DifferentiableWithinAt 𝕜 W s x) + (hW₁ : DifferentiableWithinAt 𝕜 W₁ s x) (hs : UniqueDiffWithinAt 𝕜 s x) : + lieBracketWithin 𝕜 V (W + W₁) s x = + lieBracketWithin 𝕜 V W s x + lieBracketWithin 𝕜 V W₁ s x := by + simp only [lieBracketWithin, Pi.add_apply, map_add] + rw [fderivWithin_add' hs hW hW₁, ContinuousLinearMap.add_apply] + abel + +lemma lieBracket_add_right (hW : DifferentiableAt 𝕜 W x) (hW₁ : DifferentiableAt 𝕜 W₁ x) : + lieBracket 𝕜 V (W + W₁) x = + lieBracket 𝕜 V W x + lieBracket 𝕜 V W₁ x := by + simp only [lieBracket, Pi.add_apply, map_add] + rw [fderiv_add' hW hW₁, ContinuousLinearMap.add_apply] + abel + +lemma lieBracketWithin_swap : lieBracketWithin 𝕜 V W s = - lieBracketWithin 𝕜 W V s := by + ext x; simp [lieBracketWithin] + +lemma lieBracket_swap : lieBracket 𝕜 V W x = - lieBracket 𝕜 W V x := by + simp [lieBracket] + +@[simp] lemma lieBracketWithin_self : lieBracketWithin 𝕜 V V s = 0 := by + ext x; simp [lieBracketWithin] + +@[simp] lemma lieBracket_self : lieBracket 𝕜 V V = 0 := by + ext x; simp [lieBracket] + +lemma _root_.ContDiffWithinAt.lieBracketWithin_vectorField + {m n : ℕ∞} (hV : ContDiffWithinAt 𝕜 n V s x) + (hW : ContDiffWithinAt 𝕜 n W s x) (hs : UniqueDiffOn 𝕜 s) (hmn : m + 1 ≤ n) (hx : x ∈ s) : + ContDiffWithinAt 𝕜 m (lieBracketWithin 𝕜 V W s) s x := by + apply ContDiffWithinAt.sub + · exact ContDiffWithinAt.clm_apply (hW.fderivWithin_right hs hmn hx) + (hV.of_le (le_trans le_self_add hmn)) + · exact ContDiffWithinAt.clm_apply (hV.fderivWithin_right hs hmn hx) + (hW.of_le (le_trans le_self_add hmn)) + +lemma _root_.ContDiffAt.lieBracket_vectorField {m n : ℕ∞} (hV : ContDiffAt 𝕜 n V x) + (hW : ContDiffAt 𝕜 n W x) (hmn : m + 1 ≤ n) : + ContDiffAt 𝕜 m (lieBracket 𝕜 V W) x := by + rw [← contDiffWithinAt_univ] at hV hW ⊢ + simp_rw [← lieBracketWithin_univ] + exact hV.lieBracketWithin_vectorField hW uniqueDiffOn_univ hmn (mem_univ _) + +lemma _root_.ContDiffOn.lieBracketWithin_vectorField {m n : ℕ∞} (hV : ContDiffOn 𝕜 n V s) + (hW : ContDiffOn 𝕜 n W s) (hs : UniqueDiffOn 𝕜 s) (hmn : m + 1 ≤ n) : + ContDiffOn 𝕜 m (lieBracketWithin 𝕜 V W s) s := + fun x hx ↦ (hV x hx).lieBracketWithin_vectorField (hW x hx) hs hmn hx + +lemma _root_.ContDiff.lieBracket_vectorField {m n : ℕ∞} (hV : ContDiff 𝕜 n V) + (hW : ContDiff 𝕜 n W) (hmn : m + 1 ≤ n) : + ContDiff 𝕜 m (lieBracket 𝕜 V W) := + contDiff_iff_contDiffAt.2 (fun _ ↦ hV.contDiffAt.lieBracket_vectorField hW.contDiffAt hmn) + +theorem lieBracketWithin_of_mem_nhdsWithin (st : t ∈ 𝓝[s] x) (hs : UniqueDiffWithinAt 𝕜 s x) + (hV : DifferentiableWithinAt 𝕜 V t x) (hW : DifferentiableWithinAt 𝕜 W t x) : + lieBracketWithin 𝕜 V W s x = lieBracketWithin 𝕜 V W t x := by + simp [lieBracketWithin, fderivWithin_of_mem_nhdsWithin st hs hV, + fderivWithin_of_mem_nhdsWithin st hs hW] + +theorem lieBracketWithin_subset (st : s ⊆ t) (ht : UniqueDiffWithinAt 𝕜 s x) + (hV : DifferentiableWithinAt 𝕜 V t x) (hW : DifferentiableWithinAt 𝕜 W t x) : + lieBracketWithin 𝕜 V W s x = lieBracketWithin 𝕜 V W t x := + lieBracketWithin_of_mem_nhdsWithin (nhdsWithin_mono _ st self_mem_nhdsWithin) ht hV hW + +theorem lieBracketWithin_inter (ht : t ∈ 𝓝 x) : + lieBracketWithin 𝕜 V W (s ∩ t) x = lieBracketWithin 𝕜 V W s x := by + simp [lieBracketWithin, fderivWithin_inter, ht] + +theorem lieBracketWithin_of_mem_nhds (h : s ∈ 𝓝 x) : + lieBracketWithin 𝕜 V W s x = lieBracket 𝕜 V W x := by + rw [← lieBracketWithin_univ, ← univ_inter s, lieBracketWithin_inter h] + +theorem lieBracketWithin_of_isOpen (hs : IsOpen s) (hx : x ∈ s) : + lieBracketWithin 𝕜 V W s x = lieBracket 𝕜 V W x := + lieBracketWithin_of_mem_nhds (hs.mem_nhds hx) + +theorem lieBracketWithin_eq_lieBracket (hs : UniqueDiffWithinAt 𝕜 s x) + (hV : DifferentiableAt 𝕜 V x) (hW : DifferentiableAt 𝕜 W x) : + lieBracketWithin 𝕜 V W s x = lieBracket 𝕜 V W x := by + simp [lieBracketWithin, lieBracket, fderivWithin_eq_fderiv, hs, hV, hW] + +/-- Variant of `lieBracketWithin_congr_set` where one requires the sets to coincide only in +the complement of a point. -/ +theorem lieBracketWithin_congr_set' (y : E) (h : s =ᶠ[𝓝[{y}ᶜ] x] t) : + lieBracketWithin 𝕜 V W s x = lieBracketWithin 𝕜 V W t x := by + simp [lieBracketWithin, fderivWithin_congr_set' _ h] + +theorem lieBracketWithin_congr_set (h : s =ᶠ[𝓝 x] t) : + lieBracketWithin 𝕜 V W s x = lieBracketWithin 𝕜 V W t x := + lieBracketWithin_congr_set' x <| h.filter_mono inf_le_left + +/-- Variant of `lieBracketWithin_eventually_congr_set` where one requires the sets to coincide only +in the complement of a point. -/ +theorem lieBracketWithin_eventually_congr_set' (y : E) (h : s =ᶠ[𝓝[{y}ᶜ] x] t) : + lieBracketWithin 𝕜 V W s =ᶠ[𝓝 x] lieBracketWithin 𝕜 V W t := + (eventually_nhds_nhdsWithin.2 h).mono fun _ => lieBracketWithin_congr_set' y + +theorem lieBracketWithin_eventually_congr_set (h : s =ᶠ[𝓝 x] t) : + lieBracketWithin 𝕜 V W s =ᶠ[𝓝 x] lieBracketWithin 𝕜 V W t := + lieBracketWithin_eventually_congr_set' x <| h.filter_mono inf_le_left + +theorem _root_.DifferentiableWithinAt.lieBracketWithin_congr_mono + (hV : DifferentiableWithinAt 𝕜 V s x) (hVs : EqOn V₁ V t) (hVx : V₁ x = V x) + (hW : DifferentiableWithinAt 𝕜 W s x) (hWs : EqOn W₁ W t) (hWx : W₁ x = W x) + (hxt : UniqueDiffWithinAt 𝕜 t x) (h₁ : t ⊆ s) : + lieBracketWithin 𝕜 V₁ W₁ t x = lieBracketWithin 𝕜 V W s x := by + simp [lieBracketWithin, hV.fderivWithin_congr_mono, hW.fderivWithin_congr_mono, hVs, hVx, + hWs, hWx, hxt, h₁] + +theorem _root_.Filter.EventuallyEq.lieBracketWithin_vectorField_eq + (hV : V₁ =ᶠ[𝓝[s] x] V) (hxV : V₁ x = V x) (hW : W₁ =ᶠ[𝓝[s] x] W) (hxW : W₁ x = W x) : + lieBracketWithin 𝕜 V₁ W₁ s x = lieBracketWithin 𝕜 V W s x := by + simp only [lieBracketWithin, hV.fderivWithin_eq hxV, hW.fderivWithin_eq hxW, hxV, hxW] + +theorem _root_.Filter.EventuallyEq.lieBracketWithin_vectorField_eq_of_mem + (hV : V₁ =ᶠ[𝓝[s] x] V) (hW : W₁ =ᶠ[𝓝[s] x] W) (hx : x ∈ s) : + lieBracketWithin 𝕜 V₁ W₁ s x = lieBracketWithin 𝕜 V W s x := + hV.lieBracketWithin_vectorField_eq (mem_of_mem_nhdsWithin hx hV :) + hW (mem_of_mem_nhdsWithin hx hW :) + +/-- If vector fields coincide on a neighborhood of a point within a set, then the Lie brackets +also coincide on a neighborhood of this point within this set. Version where one considers the Lie +bracket within a subset. -/ +theorem _root_.Filter.EventuallyEq.lieBracketWithin_vectorField' + (hV : V₁ =ᶠ[𝓝[s] x] V) (hW : W₁ =ᶠ[𝓝[s] x] W) (ht : t ⊆ s) : + lieBracketWithin 𝕜 V₁ W₁ t =ᶠ[𝓝[s] x] lieBracketWithin 𝕜 V W t := by + filter_upwards [hV.fderivWithin' ht (𝕜 := 𝕜), hW.fderivWithin' ht (𝕜 := 𝕜), hV, hW] + with x hV' hW' hV hW + simp [lieBracketWithin, hV', hW', hV, hW] + +protected theorem _root_.Filter.EventuallyEq.lieBracketWithin_vectorField + (hV : V₁ =ᶠ[𝓝[s] x] V) (hW : W₁ =ᶠ[𝓝[s] x] W) : + lieBracketWithin 𝕜 V₁ W₁ s =ᶠ[𝓝[s] x] lieBracketWithin 𝕜 V W s := + hV.lieBracketWithin_vectorField' hW Subset.rfl + +protected theorem _root_.Filter.EventuallyEq.lieBracketWithin_vectorField_eq_of_insert + (hV : V₁ =ᶠ[𝓝[insert x s] x] V) (hW : W₁ =ᶠ[𝓝[insert x s] x] W) : + lieBracketWithin 𝕜 V₁ W₁ s x = lieBracketWithin 𝕜 V W s x := by + apply mem_of_mem_nhdsWithin (mem_insert x s) (hV.lieBracketWithin_vectorField' hW + (subset_insert x s)) + +theorem _root_.Filter.EventuallyEq.lieBracketWithin_vectorField_eq_nhds + (hV : V₁ =ᶠ[𝓝 x] V) (hW : W₁ =ᶠ[𝓝 x] W) : + lieBracketWithin 𝕜 V₁ W₁ s x = lieBracketWithin 𝕜 V W s x := + (hV.filter_mono nhdsWithin_le_nhds).lieBracketWithin_vectorField_eq hV.self_of_nhds + (hW.filter_mono nhdsWithin_le_nhds) hW.self_of_nhds + +theorem lieBracketWithin_congr + (hV : EqOn V₁ V s) (hVx : V₁ x = V x) (hW : EqOn W₁ W s) (hWx : W₁ x = W x) : + lieBracketWithin 𝕜 V₁ W₁ s x = lieBracketWithin 𝕜 V W s x := + (hV.eventuallyEq.filter_mono inf_le_right).lieBracketWithin_vectorField_eq hVx + (hW.eventuallyEq.filter_mono inf_le_right) hWx + +/-- Version of `lieBracketWithin_congr` in which one assumes that the point belongs to the +given set. -/ +theorem lieBracketWithin_congr' (hV : EqOn V₁ V s) (hW : EqOn W₁ W s) (hx : x ∈ s) : + lieBracketWithin 𝕜 V₁ W₁ s x = lieBracketWithin 𝕜 V W s x := + lieBracketWithin_congr hV (hV hx) hW (hW hx) + +theorem _root_.Filter.EventuallyEq.lieBracket_vectorField_eq + (hV : V₁ =ᶠ[𝓝 x] V) (hW : W₁ =ᶠ[𝓝 x] W) : + lieBracket 𝕜 V₁ W₁ x = lieBracket 𝕜 V W x := by + rw [← lieBracketWithin_univ, ← lieBracketWithin_univ, hV.lieBracketWithin_vectorField_eq_nhds hW] + +protected theorem _root_.Filter.EventuallyEq.lieBracket_vectorField + (hV : V₁ =ᶠ[𝓝 x] V) (hW : W₁ =ᶠ[𝓝 x] W) : lieBracket 𝕜 V₁ W₁ =ᶠ[𝓝 x] lieBracket 𝕜 V W := by + filter_upwards [hV.eventuallyEq_nhds, hW.eventuallyEq_nhds] with y hVy hWy + exact hVy.lieBracket_vectorField_eq hWy + +/-- The Lie bracket of vector fields in vector spaces satisfies the Leibniz identity +`[U, [V, W]] = [[U, V], W] + [V, [U, W]]`. -/ +lemma leibniz_identity_lieBracketWithin_of_isSymmSndFDerivWithinAt + {U V W : E → E} {s : Set E} {x : E} (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) + (hU : ContDiffWithinAt 𝕜 2 U s x) (hV : ContDiffWithinAt 𝕜 2 V s x) + (hW : ContDiffWithinAt 𝕜 2 W s x) + (h'U : IsSymmSndFDerivWithinAt 𝕜 U s x) (h'V : IsSymmSndFDerivWithinAt 𝕜 V s x) + (h'W : IsSymmSndFDerivWithinAt 𝕜 W s x) : + lieBracketWithin 𝕜 U (lieBracketWithin 𝕜 V W s) s x = + lieBracketWithin 𝕜 (lieBracketWithin 𝕜 U V s) W s x + + lieBracketWithin 𝕜 V (lieBracketWithin 𝕜 U W s) s x := by + simp only [lieBracketWithin_eq, map_sub] + have aux₁ {U V : E → E} (hU : ContDiffWithinAt 𝕜 2 U s x) (hV : ContDiffWithinAt 𝕜 2 V s x) : + DifferentiableWithinAt 𝕜 (fun x ↦ (fderivWithin 𝕜 V s x) (U x)) s x := + have := hV.fderivWithin_right_apply (hU.of_le one_le_two) hs le_rfl hx + this.differentiableWithinAt le_rfl + have aux₂ {U V : E → E} (hU : ContDiffWithinAt 𝕜 2 U s x) (hV : ContDiffWithinAt 𝕜 2 V s x) : + fderivWithin 𝕜 (fun y ↦ (fderivWithin 𝕜 U s y) (V y)) s x = + (fderivWithin 𝕜 U s x).comp (fderivWithin 𝕜 V s x) + + (fderivWithin 𝕜 (fderivWithin 𝕜 U s) s x).flip (V x) := by + refine fderivWithin_clm_apply (hs x hx) ?_ (hV.differentiableWithinAt one_le_two) + exact (hU.fderivWithin_right hs le_rfl hx).differentiableWithinAt le_rfl + rw [fderivWithin_sub (hs x hx) (aux₁ hV hW) (aux₁ hW hV)] + rw [fderivWithin_sub (hs x hx) (aux₁ hU hV) (aux₁ hV hU)] + rw [fderivWithin_sub (hs x hx) (aux₁ hU hW) (aux₁ hW hU)] + rw [aux₂ hW hV, aux₂ hV hW, aux₂ hV hU, aux₂ hU hV, aux₂ hW hU, aux₂ hU hW] + simp only [ContinuousLinearMap.coe_sub', Pi.sub_apply, ContinuousLinearMap.add_apply, + ContinuousLinearMap.coe_comp', Function.comp_apply, ContinuousLinearMap.flip_apply, h'V.eq, + h'U.eq, h'W.eq] + abel + +/-- The Lie bracket of vector fields in vector spaces satisfies the Leibniz identity +`[U, [V, W]] = [[U, V], W] + [V, [U, W]]`. -/ +lemma leibniz_identity_lieBracketWithin [IsRCLikeNormedField 𝕜] {U V W : E → E} {s : Set E} {x : E} + (hs : UniqueDiffOn 𝕜 s) (h'x : x ∈ closure (interior s)) (hx : x ∈ s) + (hU : ContDiffWithinAt 𝕜 2 U s x) (hV : ContDiffWithinAt 𝕜 2 V s x) + (hW : ContDiffWithinAt 𝕜 2 W s x) : + lieBracketWithin 𝕜 U (lieBracketWithin 𝕜 V W s) s x = + lieBracketWithin 𝕜 (lieBracketWithin 𝕜 U V s) W s x + + lieBracketWithin 𝕜 V (lieBracketWithin 𝕜 U W s) s x := by + apply leibniz_identity_lieBracketWithin_of_isSymmSndFDerivWithinAt hs hx hU hV hW + · exact hU.isSymmSndFDerivWithinAt le_rfl hs h'x hx + · exact hV.isSymmSndFDerivWithinAt le_rfl hs h'x hx + · exact hW.isSymmSndFDerivWithinAt le_rfl hs h'x hx + +/-- The Lie bracket of vector fields in vector spaces satisfies the Leibniz identity +`[U, [V, W]] = [[U, V], W] + [V, [U, W]]`. -/ +lemma leibniz_identity_lieBracket [IsRCLikeNormedField 𝕜] {U V W : E → E} {x : E} + (hU : ContDiffAt 𝕜 2 U x) (hV : ContDiffAt 𝕜 2 V x) (hW : ContDiffAt 𝕜 2 W x) : + lieBracket 𝕜 U (lieBracket 𝕜 V W) x = + lieBracket 𝕜 (lieBracket 𝕜 U V) W x + lieBracket 𝕜 V (lieBracket 𝕜 U W) x := by + simp only [← lieBracketWithin_univ, ← contDiffWithinAt_univ] at hU hV hW ⊢ + exact leibniz_identity_lieBracketWithin uniqueDiffOn_univ (by simp) (mem_univ _) hU hV hW + +end VectorField From 92f63d5c1988987aa375830b9b41cbfd269c3625 Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 20 Nov 2024 14:53:33 +0000 Subject: [PATCH 154/829] fix(CI): fail if tests are noisy (#19268) [Zulip discussion](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/CI.3A.20noisy.20.22test.20mathlib.22/near/483367292) #19269 adds a test file with `#eval ""` and tests fail on that PR. --- .github/build.in.yml | 2 +- .github/workflows/bors.yml | 2 +- .github/workflows/build.yml | 2 +- .github/workflows/build_fork.yml | 2 +- 4 files changed, 4 insertions(+), 4 deletions(-) diff --git a/.github/build.in.yml b/.github/build.in.yml index cf758b28c47f3..bab96d693f6b2 100644 --- a/.github/build.in.yml +++ b/.github/build.in.yml @@ -270,7 +270,7 @@ jobs: with: linters: gcc run: - lake test + lake --iofail test - name: check for unused imports id: shake diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index 68c1d03a1f5f2..e59f48d2c7267 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -280,7 +280,7 @@ jobs: with: linters: gcc run: - lake test + lake --iofail test - name: check for unused imports id: shake diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 7edf488cd7923..1c778a5720c23 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -287,7 +287,7 @@ jobs: with: linters: gcc run: - lake test + lake --iofail test - name: check for unused imports id: shake diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index 76448210f9cb0..e92f87fd50a68 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -284,7 +284,7 @@ jobs: with: linters: gcc run: - lake test + lake --iofail test - name: check for unused imports id: shake From 238cb187bb518855c7150ac44faa919685eb682b Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 20 Nov 2024 15:32:16 +0000 Subject: [PATCH 155/829] feat(CI): merge bors (#19078) Unify the three `bors d`, `bors merge` and `bord r+` actions into a single one. Prefer using `TRIAGE` token for GH API quota. --- .github/workflows/maintainer_bors.yml | 93 +++++++++++++++++++++++++++ 1 file changed, 93 insertions(+) create mode 100644 .github/workflows/maintainer_bors.yml diff --git a/.github/workflows/maintainer_bors.yml b/.github/workflows/maintainer_bors.yml new file mode 100644 index 0000000000000..d81dded626572 --- /dev/null +++ b/.github/workflows/maintainer_bors.yml @@ -0,0 +1,93 @@ +name: Add "ready-to-merge" and "delegated" label + +# triggers the action when +on: + # the PR receives a comment + issue_comment: + types: [created] + # the PR receives a review + pull_request_review: + # whether or not it is accompanied by a comment + types: [submitted] + # the PR receives a review comment + pull_request_review_comment: + types: [created] + +jobs: + add_ready_to_merge_label: + # we set some variables. The ones of the form `${{ X }}${{ Y }}` are typically not + # both set simultaneously: depending on the event that triggers the PR, usually only one is set + env: + AUTHOR: ${{ github.event.comment.user.login }}${{ github.event.review.user.login }} + PR_NUMBER: ${{ github.event.issue.number }}${{ github.event.pull_request.number }} + COMMENT_EVENT: ${{ github.event.comment.body }} + COMMENT_REVIEW: ${{ github.event.review.body }} + PR_TITLE_ISSUE: ${{ github.event.issue.title }} + PR_TITLE_PR: ${{ github.event.pull_request.title }} + PR_URL: ${{ github.event.issue.html_url }}${{ github.event.pull_request.html_url }} + EVENT_NAME: ${{ github.event_name }} + name: Add ready-to-merge or delegated label + runs-on: ubuntu-latest + steps: + - name: Find bors merge/delegate + id: merge_or_delegate + run: | + COMMENT="${COMMENT_EVENT}${COMMENT_REVIEW}" + # we strip `\r` since line endings from GitHub contain this character + COMMENT="${COMMENT//$'\r'/}" + # for debugging, we print some information + printf '%s' "${COMMENT}" | hexdump -cC + printf 'Comment:"%s"\n' "${COMMENT}" + m_or_d="$(printf '%s' "${COMMENT}" | + sed -n 's=^bors *\(merge\|r+\) *$=ready-to-merge=p; s=^bors *d.*=delegated=p' | head -1)" + + printf $'"bors delegate" or "bors merge" found? \'%s\'\n' "${m_or_d}" + + printf $'mOrD=%s\n' "${m_or_d}" > "${GITHUB_OUTPUT}" + if [ "${AUTHOR}" == 'leanprover-community-mathlib4-bot' ] || + [ "${AUTHOR}" == 'leanprover-community-bot-assistant' ] + then + printf $'bot=true\n' + printf $'bot=true\n' > "${GITHUB_OUTPUT}" + else + printf $'bot=false\n' + printf $'bot=false\n' > "${GITHUB_OUTPUT}" + fi + + - name: Check whether user is a mathlib admin + id: user_permission + if: ${{ ! steps.merge_or_delegate.outputs.mOrD == '' }} + uses: actions-cool/check-user-permission@v2 + with: + require: 'admin' + # review(_comment) use + # require: 'write' + # token: ${{ secrets.TRIAGE_TOKEN }} + + - name: Add ready-to-merge or delegated label + id: add_label + if: ${{ ! steps.merge_or_delegate.outputs.mOrD == '' && + ( steps.user_permission.outputs.require-result == 'true' || + steps.merge_or_delegate.outputs.bot == 'true' ) }} + uses: octokit/request-action@v2.x + with: + # check is this ok? was /repos/:repository/issues/:issue_number/labels + route: POST /repos/:repository/issues/${PR_NUMBER}/labels + repository: ${{ github.repository }} + issue_number: ${PR_NUMBER} + labels: '["${{ steps.merge_or_delegate.outputs.mOrD }}"]' + env: + GITHUB_TOKEN: ${{ secrets.TRIAGE_TOKEN }} # comment uses ${{ secrets.GITHUB_TOKEN }} + + - if: ${{ ! steps.merge_or_delegate.outputs.mOrD == '' && + ( steps.user_permission.outputs.require-result == 'true' || + steps.merge_or_delegate.outputs.bot == 'true' ) }} + id: remove_labels + name: Remove "awaiting-author" + # we use curl rather than octokit/request-action so that the job won't fail + # (and send an annoying email) if the labels don't exist + run: | + curl --request DELETE \ + --url "https://api.github.com/repos/${{ github.repository }}/issues/${PR_NUMBER}/labels/awaiting-author" \ + --header 'authorization: Bearer ${{ secrets.TRIAGE_TOKEN }}' + # comment uses ${{ secrets.GITHUB_TOKEN }} From 8eed8831892ac64dca2e5d290fdbc2dcfeecfd83 Mon Sep 17 00:00:00 2001 From: sven-manthe Date: Wed, 20 Nov 2024 16:52:58 +0000 Subject: [PATCH 156/829] chore(CategoryTheory/IsConnected): add refl etc annotations (#19017) --- Mathlib/CategoryTheory/IsConnected.lean | 11 ++++++----- 1 file changed, 6 insertions(+), 5 deletions(-) diff --git a/Mathlib/CategoryTheory/IsConnected.lean b/Mathlib/CategoryTheory/IsConnected.lean index 96bb16ce5927a..f4fa03b9c28a6 100644 --- a/Mathlib/CategoryTheory/IsConnected.lean +++ b/Mathlib/CategoryTheory/IsConnected.lean @@ -263,11 +263,11 @@ theorem isConnected_of_isConnected_op [IsConnected Jᵒᵖ] : IsConnected J := def Zag (j₁ j₂ : J) : Prop := Nonempty (j₁ ⟶ j₂) ∨ Nonempty (j₂ ⟶ j₁) -theorem Zag.refl (X : J) : Zag X X := Or.inl ⟨𝟙 _⟩ +@[refl] theorem Zag.refl (X : J) : Zag X X := Or.inl ⟨𝟙 _⟩ theorem zag_symmetric : Symmetric (@Zag J _) := fun _ _ h => h.symm -theorem Zag.symm {j₁ j₂ : J} (h : Zag j₁ j₂) : Zag j₂ j₁ := zag_symmetric h +@[symm] theorem Zag.symm {j₁ j₂ : J} (h : Zag j₁ j₂) : Zag j₂ j₁ := zag_symmetric h theorem Zag.of_hom {j₁ j₂ : J} (f : j₁ ⟶ j₂) : Zag j₁ j₂ := Or.inl ⟨f⟩ @@ -286,11 +286,12 @@ theorem zigzag_equivalence : _root_.Equivalence (@Zigzag J _) := _root_.Equivalence.mk Relation.reflexive_reflTransGen (fun h => zigzag_symmetric h) (fun h g => Relation.transitive_reflTransGen h g) -theorem Zigzag.refl (X : J) : Zigzag X X := zigzag_equivalence.refl _ +@[refl] theorem Zigzag.refl (X : J) : Zigzag X X := zigzag_equivalence.refl _ -theorem Zigzag.symm {j₁ j₂ : J} (h : Zigzag j₁ j₂) : Zigzag j₂ j₁ := zigzag_symmetric h +@[symm] theorem Zigzag.symm {j₁ j₂ : J} (h : Zigzag j₁ j₂) : Zigzag j₂ j₁ := zigzag_symmetric h -theorem Zigzag.trans {j₁ j₂ j₃ : J} (h₁ : Zigzag j₁ j₂) (h₂ : Zigzag j₂ j₃) : Zigzag j₁ j₃ := +@[trans] theorem Zigzag.trans {j₁ j₂ j₃ : J} (h₁ : Zigzag j₁ j₂) (h₂ : Zigzag j₂ j₃) : + Zigzag j₁ j₃ := zigzag_equivalence.trans h₁ h₂ theorem Zigzag.of_zag {j₁ j₂ : J} (h : Zag j₁ j₂) : Zigzag j₁ j₂ := From 46d221cdeabdc58a33e8724bd000be0f730d31ff Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Wed, 20 Nov 2024 17:06:21 +0000 Subject: [PATCH 157/829] feat(NumberTheory/LSeries/DirichletContinuation): results on logarithmic derivatives (#19254) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit As a further step in the direction of Dirichlet's Theorem on primes in AP, this PR adds results on (negative) logarithmic derivatives of `L χ` for a non-trivial Dirichlet character `χ` and of `s ↦ (L χ s) * (s - 1)` when `χ` is trivial. See [here on Zulip](https://leanprover.zulipchat.com/#narrow/channel/144837-PR-reviews/topic/Prerequisites.20for.20PNT.20and.20Dirichlet's.20Thm/near/482543603). Co-authored-by: Michael Stoll <99838730+MichaelStollBayreuth@users.noreply.github.com> --- .../LSeries/DirichletContinuation.lean | 91 ++++++++++++++++++- 1 file changed, 88 insertions(+), 3 deletions(-) diff --git a/Mathlib/NumberTheory/LSeries/DirichletContinuation.lean b/Mathlib/NumberTheory/LSeries/DirichletContinuation.lean index e4b323edd1d75..033ad5dde6fc6 100644 --- a/Mathlib/NumberTheory/LSeries/DirichletContinuation.lean +++ b/Mathlib/NumberTheory/LSeries/DirichletContinuation.lean @@ -84,11 +84,13 @@ lemma deriv_LFunction_eq_deriv_LSeries (χ : DirichletCharacter ℂ N) {s : ℂ} The L-function of a Dirichlet character is differentiable, except at `s = 1` if the character is trivial. -/ +@[fun_prop] lemma differentiableAt_LFunction (χ : DirichletCharacter ℂ N) (s : ℂ) (hs : s ≠ 1 ∨ χ ≠ 1) : DifferentiableAt ℂ (LFunction χ) s := ZMod.differentiableAt_LFunction χ s (hs.imp_right χ.sum_eq_zero_of_ne_one) /-- The L-function of a non-trivial Dirichlet character is differentiable everywhere. -/ +@[fun_prop] lemma differentiable_LFunction {χ : DirichletCharacter ℂ N} (hχ : χ ≠ 1) : Differentiable ℂ (LFunction χ) := (differentiableAt_LFunction _ · <| Or.inr hχ) @@ -113,7 +115,7 @@ lemma Even.LFunction_neg_two_mul_nat {χ : DirichletCharacter ℂ N} (hχ : Even ZMod.LFunction_neg_two_mul_nat_sub_one hχ.to_fun n /-! -## Results on changing levels +### Results on changing levels -/ private lemma LFunction_changeLevel_aux {M N : ℕ} [NeZero M] [NeZero N] (hMN : M ∣ N) @@ -156,7 +158,7 @@ lemma LFunction_changeLevel {M N : ℕ} [NeZero M] [NeZero N] (hMN : M ∣ N) · exact LFunction_changeLevel_aux hMN χ h /-! -## The `L`-function of the trivial character mod `N` +### The `L`-function of the trivial character mod `N` -/ /-- The `L`-function of the trivial character mod `N`. -/ @@ -189,7 +191,7 @@ lemma LFunctionTrivChar_residue_one : fun_prop /-! -## Completed L-functions and the functional equation +### Completed L-functions and the functional equation -/ section gammaFactor @@ -307,3 +309,86 @@ theorem completedLFunction_one_sub {χ : DirichletCharacter ℂ N} (hχ : IsPrim end IsPrimitive end DirichletCharacter + +/-! +### The logarithmic derivative of the L-function of a Dirichlet character + +We show that `s ↦ -(L' χ s) / L χ s + 1 / (s - 1)` is continuous outside the zeros of `L χ` +when `χ` is a trivial Dirichlet character and that `-L' χ / L χ` is continuous outside +the zeros of `L χ` when `χ` is nontrivial. +-/ + +namespace DirichletCharacter + +open Complex + +section trivial + +variable (n : ℕ) [NeZero n] + +/-- The function obtained by "multiplying away" the pole of `L χ` for a trivial Dirichlet +character `χ`. Its (negative) logarithmic derivative is used to prove Dirichlet's Theorem +on primes in arithmetic progression. -/ +noncomputable abbrev LFunctionTrivChar₁ : ℂ → ℂ := + Function.update (fun s ↦ (s - 1) * LFunctionTrivChar n s) 1 + (∏ p ∈ n.primeFactors, (1 - (p : ℂ)⁻¹)) + +lemma LFunctionTrivChar₁_apply_one_ne_zero : LFunctionTrivChar₁ n 1 ≠ 0 := by + simp only [Function.update_same] + refine Finset.prod_ne_zero_iff.mpr fun p hp ↦ ?_ + simpa only [ne_eq, sub_ne_zero, one_eq_inv, Nat.cast_eq_one] + using (Nat.prime_of_mem_primeFactors hp).ne_one + +/-- `s ↦ (s - 1) * L χ s` is an entire function when `χ` is a trivial Dirichlet character. -/ +lemma differentiable_LFunctionTrivChar₁ : Differentiable ℂ (LFunctionTrivChar₁ n) := by + rw [← differentiableOn_univ, + ← differentiableOn_compl_singleton_and_continuousAt_iff (c := 1) Filter.univ_mem] + refine ⟨DifferentiableOn.congr (f := fun s ↦ (s - 1) * LFunctionTrivChar n s) + (fun _ hs ↦ DifferentiableAt.differentiableWithinAt <| by fun_prop (disch := simp_all [hs])) + fun _ hs ↦ Function.update_noteq (Set.mem_diff_singleton.mp hs).2 .., + continuousWithinAt_compl_self.mp ?_⟩ + simpa only [continuousWithinAt_compl_self, continuousAt_update_same] + using LFunctionTrivChar_residue_one + +lemma deriv_LFunctionTrivChar₁_apply_of_ne_one {s : ℂ} (hs : s ≠ 1) : + deriv (LFunctionTrivChar₁ n) s = + (s - 1) * deriv (LFunctionTrivChar n) s + LFunctionTrivChar n s := by + have H : deriv (LFunctionTrivChar₁ n) s = + deriv (fun w ↦ (w - 1) * LFunctionTrivChar n w) s := by + refine eventuallyEq_iff_exists_mem.mpr ?_ |>.deriv_eq + exact ⟨_, isOpen_ne.mem_nhds hs, fun _ hw ↦ Function.update_noteq (Set.mem_setOf.mp hw) ..⟩ + rw [H, deriv_mul (by fun_prop) (differentiableAt_LFunction _ s (.inl hs)), deriv_sub_const, + deriv_id'', one_mul, add_comm] + +/-- The negative logarithmtic derivative of `s ↦ (s - 1) * L χ s` for a trivial +Dirichlet character `χ` is continuous away from the zeros of `L χ` (including at `s = 1`). -/ +lemma continuousOn_neg_logDeriv_LFunctionTrivChar₁ : + ContinuousOn (fun s ↦ -deriv (LFunctionTrivChar₁ n) s / LFunctionTrivChar₁ n s) + {s | s = 1 ∨ LFunctionTrivChar n s ≠ 0} := by + simp_rw [neg_div] + have h := differentiable_LFunctionTrivChar₁ n + refine ((h.contDiff.continuous_deriv le_rfl).continuousOn.div + h.continuous.continuousOn fun w hw ↦ ?_).neg + rcases eq_or_ne w 1 with rfl | hw' + · exact LFunctionTrivChar₁_apply_one_ne_zero _ + · rw [LFunctionTrivChar₁, Function.update_noteq hw', mul_ne_zero_iff] + exact ⟨sub_ne_zero_of_ne hw', (Set.mem_setOf.mp hw).resolve_left hw'⟩ + +end trivial + +section nontrivial + +variable {n : ℕ} [NeZero n] {χ : DirichletCharacter ℂ n} + +/-- The negative logarithmic derivative of the L-function of a nontrivial Dirichlet character +is continuous away from the zeros of the L-function. -/ +lemma continuousOn_neg_logDeriv_LFunction_of_nontriv (hχ : χ ≠ 1) : + ContinuousOn (fun s ↦ -deriv (LFunction χ) s / LFunction χ s) {s | LFunction χ s ≠ 0} := by + simp only [neg_div] + have h := differentiable_LFunction hχ + exact ((h.contDiff.continuous_deriv le_rfl).continuousOn.div + h.continuous.continuousOn fun _ hw ↦ hw).neg + +end nontrivial + +end DirichletCharacter From a5d04c838085d6c3554d59278c04126b62314757 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Wed, 20 Nov 2024 17:19:12 +0000 Subject: [PATCH 158/829] chore(RingTheory): improve and generalize submodule localization API (#19118) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In Algebra/Module/LocalizedModule/Submodule, we introduce `Submodule.localized₀`, which is localization of a submodule considered as a submodule over the base ring rather than the localization. This allows us to talk about the localization of a submodule without choosing a specific ring localization, just like what `IsLocalizedModule` allows us to do for localization of a module. As applications, `Rₚ` and every hypothesis depending on it are completely removed from the statement of `Submodule.le_of_localization_maximal`, etc. in RingTheory/LocalProperties/Submodule. We also reorder lemmas in this file to make the development more natural and golf the proofs (making use of `Module.eqIdeal` introduced in RingTheory/Ideal/Defs; also add some trivial lemmas. In RingTheory/LocalProperties/Projective, we fix a proof due to removed explicit argument `Rₚ` and remove some unnecessary lines. --- .../Module/LocalizedModule/Submodule.lean | 94 ++++++++---- Mathlib/RingTheory/Ideal/Defs.lean | 9 ++ .../LocalProperties/Projective.lean | 18 +-- .../RingTheory/LocalProperties/Submodule.lean | 139 +++++++++++------- 4 files changed, 163 insertions(+), 97 deletions(-) diff --git a/Mathlib/Algebra/Module/LocalizedModule/Submodule.lean b/Mathlib/Algebra/Module/LocalizedModule/Submodule.lean index 786335636ddb0..650844736f7ca 100644 --- a/Mathlib/Algebra/Module/LocalizedModule/Submodule.lean +++ b/Mathlib/Algebra/Module/LocalizedModule/Submodule.lean @@ -28,49 +28,73 @@ Results about localizations of submodules and quotient modules are provided in t open nonZeroDivisors -universe u u' v v' w w' - -variable {R : Type u} (S : Type u') {M : Type v} {N : Type v'} -variable [CommRing R] [CommRing S] [AddCommGroup M] [AddCommGroup N] +variable {R S M N : Type*} +variable (S) [CommSemiring R] [CommSemiring S] [AddCommMonoid M] [AddCommMonoid N] variable [Module R M] [Module R N] [Algebra R S] [Module S N] [IsScalarTower R S N] variable (p : Submonoid R) [IsLocalization p S] (f : M →ₗ[R] N) [IsLocalizedModule p f] variable (M' : Submodule R M) -/-- Let `S` be the localization of `R` at `p` and `N` be the localization of `M` at `p`. -This is the localization of an `R`-submodule of `M` viewed as an `S`-submodule of `N`. -/ -def Submodule.localized' : Submodule S N where +namespace Submodule + +/-- Let `N` be a localization of an `R`-module `M` at `p`. +This is the localization of an `R`-submodule of `M` viewed as an `R`-submodule of `N`. -/ +def localized₀ : Submodule R N where carrier := { x | ∃ m ∈ M', ∃ s : p, IsLocalizedModule.mk' f m s = x } - add_mem' := fun {x} {y} ⟨m, hm, s, hx⟩ ⟨n, hn, t, hy⟩ ↦ ⟨t • m + s • n, add_mem (M'.smul_mem t hm) + add_mem' := fun {x y} ⟨m, hm, s, hx⟩ ⟨n, hn, t, hy⟩ ↦ ⟨t • m + s • n, add_mem (M'.smul_mem t hm) (M'.smul_mem s hn), s * t, by rw [← hx, ← hy, IsLocalizedModule.mk'_add_mk']⟩ zero_mem' := ⟨0, zero_mem _, 1, by simp⟩ - smul_mem' := fun r x h ↦ by - have ⟨m, hm, s, hx⟩ := h + smul_mem' r x := by + rintro ⟨m, hm, s, hx⟩ + exact ⟨r • m, smul_mem M' _ hm, s, by rw [IsLocalizedModule.mk'_smul, hx]⟩ + +/-- Let `S` be the localization of `R` at `p` and `N` be a localization of `M` at `p`. +This is the localization of an `R`-submodule of `M` viewed as an `S`-submodule of `N`. -/ +def localized' : Submodule S N where + __ := localized₀ p f M' + smul_mem' := fun r x ⟨m, hm, s, hx⟩ ↦ by have ⟨y, t, hyt⟩ := IsLocalization.mk'_surjective p r exact ⟨y • m, M'.smul_mem y hm, t * s, by simp [← hyt, ← hx, IsLocalizedModule.mk'_smul_mk']⟩ -lemma Submodule.mem_localized' (x : N) : - x ∈ Submodule.localized' S p f M' ↔ ∃ m ∈ M', ∃ s : p, IsLocalizedModule.mk' f m s = x := +lemma mem_localized₀ (x : N) : + x ∈ localized₀ p f M' ↔ ∃ m ∈ M', ∃ s : p, IsLocalizedModule.mk' f m s = x := + Iff.rfl + +lemma mem_localized' (x : N) : + x ∈ localized' S p f M' ↔ ∃ m ∈ M', ∃ s : p, IsLocalizedModule.mk' f m s = x := Iff.rfl +/-- `localized₀` is the same as `localized'` considered as a submodule over the base ring. -/ +lemma restrictScalars_localized' : + (localized' S p f M').restrictScalars R = localized₀ p f M' := + rfl + /-- The localization of an `R`-submodule of `M` at `p` viewed as an `Rₚ`-submodule of `Mₚ`. -/ -abbrev Submodule.localized : Submodule (Localization p) (LocalizedModule p M) := +abbrev localized : Submodule (Localization p) (LocalizedModule p M) := M'.localized' (Localization p) p (LocalizedModule.mkLinearMap p M) @[simp] -lemma Submodule.localized'_bot : (⊥ : Submodule R M).localized' S p f = ⊥ := by +lemma localized₀_bot : (⊥ : Submodule R M).localized₀ p f = ⊥ := by rw [← le_bot_iff] rintro _ ⟨_, rfl, s, rfl⟩ simp only [IsLocalizedModule.mk'_zero, mem_bot] @[simp] -lemma Submodule.localized'_top : (⊤ : Submodule R M).localized' S p f = ⊤ := by +lemma localized'_bot : (⊥ : Submodule R M).localized' S p f = ⊥ := + SetLike.ext' (by apply SetLike.ext'_iff.mp <| Submodule.localized₀_bot p f) + +@[simp] +lemma localized₀_top : (⊤ : Submodule R M).localized₀ p f = ⊤ := by rw [← top_le_iff] rintro x _ obtain ⟨⟨x, s⟩, rfl⟩ := IsLocalizedModule.mk'_surjective p f x exact ⟨x, trivial, s, rfl⟩ @[simp] -lemma Submodule.localized'_span (s : Set M) : (span R s).localized' S p f = span S (f '' s) := by +lemma localized'_top : (⊤ : Submodule R M).localized' S p f = ⊤ := + SetLike.ext' (by apply SetLike.ext'_iff.mp <| Submodule.localized₀_top p f) + +@[simp] +lemma localized'_span (s : Set M) : (span R s).localized' S p f = span S (f '' s) := by apply le_antisymm · rintro _ ⟨x, hx, t, rfl⟩ have := IsLocalizedModule.mk'_smul_mk' S f 1 x t 1 @@ -87,29 +111,45 @@ lemma Submodule.localized'_span (s : Set M) : (span R s).localized' S p f = span /-- The localization map of a submodule. -/ @[simps!] -def Submodule.toLocalized' : M' →ₗ[R] M'.localized' S p f := - f.restrict (q := (M'.localized' S p f).restrictScalars R) (fun x hx ↦ ⟨x, hx, 1, by simp⟩) +def toLocalized₀ : M' →ₗ[R] M'.localized₀ p f := f.restrict fun x hx ↦ ⟨x, hx, 1, by simp⟩ /-- The localization map of a submodule. -/ -abbrev Submodule.toLocalized : M' →ₗ[R] M'.localized p := +@[simps!] +def toLocalized' : M' →ₗ[R] M'.localized' S p f := toLocalized₀ p f M' + +/-- The localization map of a submodule. -/ +abbrev toLocalized : M' →ₗ[R] M'.localized p := M'.toLocalized' (Localization p) p (LocalizedModule.mkLinearMap p M) -instance Submodule.isLocalizedModule : IsLocalizedModule p (M'.toLocalized' S p f) where +instance : IsLocalizedModule p (M'.toLocalized₀ p f) where map_units x := by simp_rw [Module.End_isUnit_iff] constructor · exact fun _ _ e ↦ Subtype.ext (IsLocalizedModule.smul_injective f x (congr_arg Subtype.val e)) - · rintro m - use (IsLocalization.mk' S 1 x) • m - rw [Module.algebraMap_end_apply, ← smul_assoc, IsLocalization.smul_mk'_one, - IsLocalization.mk'_self', one_smul] + · rintro ⟨_, m, hm, s, rfl⟩ + refine ⟨⟨IsLocalizedModule.mk' f m (s * x), ⟨_, hm, _, rfl⟩⟩, Subtype.ext ?_⟩ + rw [Module.algebraMap_end_apply, SetLike.val_smul_of_tower, + ← IsLocalizedModule.mk'_smul, ← Submonoid.smul_def, IsLocalizedModule.mk'_cancel_right] surj' := by rintro ⟨y, x, hx, s, rfl⟩ exact ⟨⟨⟨x, hx⟩, s⟩, by ext; simp⟩ exists_of_eq e := by simpa [Subtype.ext_iff] using IsLocalizedModule.exists_of_eq (S := p) (f := f) (congr_arg Subtype.val e) +instance isLocalizedModule : IsLocalizedModule p (M'.toLocalized' S p f) := + inferInstanceAs (IsLocalizedModule p (M'.toLocalized₀ p f)) + +end Submodule + +section Quotient + +variable {R S M N : Type*} +variable (S) [CommRing R] [CommRing S] [AddCommGroup M] [AddCommGroup N] +variable [Module R M] [Module R N] [Algebra R S] [Module S N] [IsScalarTower R S N] +variable (p : Submonoid R) [IsLocalization p S] (f : M →ₗ[R] N) [IsLocalizedModule p f] +variable (M' : Submodule R M) + /-- The localization map of a quotient module. -/ def Submodule.toLocalizedQuotient' : M ⧸ M' →ₗ[R] N ⧸ M'.localized' S p f := Submodule.mapQ M' ((M'.localized' S p f).restrictScalars R) f (fun x hx ↦ ⟨x, hx, 1, by simp⟩) @@ -147,10 +187,12 @@ instance IsLocalizedModule.toLocalizedQuotient' (M' : Submodule R M) : instance (M' : Submodule R M) : IsLocalizedModule p (M'.toLocalizedQuotient p) := IsLocalizedModule.toLocalizedQuotient' _ _ _ _ +end Quotient + section LinearMap -variable {P : Type w} [AddCommGroup P] [Module R P] -variable {Q : Type w'} [AddCommGroup Q] [Module R Q] [Module S Q] [IsScalarTower R S Q] +variable {P : Type*} [AddCommGroup P] [Module R P] +variable {Q : Type*} [AddCommGroup Q] [Module R Q] [Module S Q] [IsScalarTower R S Q] variable (f' : P →ₗ[R] Q) [IsLocalizedModule p f'] lemma LinearMap.localized'_ker_eq_ker_localizedMap (g : M →ₗ[R] P) : diff --git a/Mathlib/RingTheory/Ideal/Defs.lean b/Mathlib/RingTheory/Ideal/Defs.lean index d2eb82b15fdd4..ba78f8a4b205b 100644 --- a/Mathlib/RingTheory/Ideal/Defs.lean +++ b/Mathlib/RingTheory/Ideal/Defs.lean @@ -68,6 +68,15 @@ theorem unit_mul_mem_iff_mem {x y : α} (hy : IsUnit y) : y * x ∈ I ↔ x ∈ end Ideal +/-- For two elements `m` and `m'` in an `R`-module `M`, the set of elements `r : R` with +equal scalar product with `m` and `m'` is an ideal of `R`. If `M` is a group, this coincides +with the kernel of `LinearMap.toSpanSingleton R M (m - m')`. -/ +def Module.eqIdeal (R) {M} [Semiring R] [AddCommMonoid M] [Module R M] (m m' : M) : Ideal R where + carrier := {r : R | r • m = r • m'} + add_mem' h h' := by simpa [add_smul] using congr($h + $h') + zero_mem' := by simp_rw [Set.mem_setOf, zero_smul] + smul_mem' _ _ h := by simpa [mul_smul] using congr(_ • $h) + end Semiring section CommSemiring diff --git a/Mathlib/RingTheory/LocalProperties/Projective.lean b/Mathlib/RingTheory/LocalProperties/Projective.lean index d6a190e4ac15f..a2c4615a7942e 100644 --- a/Mathlib/RingTheory/LocalProperties/Projective.lean +++ b/Mathlib/RingTheory/LocalProperties/Projective.lean @@ -31,8 +31,8 @@ variable [AddCommGroup N'] [Module R N'] (S : Submonoid R) theorem Module.projective_of_isLocalizedModule {Rₛ Mₛ} [AddCommGroup Mₛ] [Module R Mₛ] [CommRing Rₛ] [Algebra R Rₛ] [Module Rₛ Mₛ] [IsScalarTower R Rₛ Mₛ] (S) (f : M →ₗ[R] Mₛ) [IsLocalization S Rₛ] [IsLocalizedModule S f] [Module.Projective R M] : - Module.Projective Rₛ Mₛ := - Projective.of_equiv (IsLocalizedModule.isBaseChange S Rₛ f).equiv + Module.Projective Rₛ Mₛ := + Projective.of_equiv (IsLocalizedModule.isBaseChange S Rₛ f).equiv theorem LinearMap.split_surjective_of_localization_maximal (f : M →ₗ[R] N) [Module.FinitePresentation R N] @@ -41,18 +41,8 @@ theorem LinearMap.split_surjective_of_localization_maximal (LocalizedModule.map I.primeCompl f).comp g = LinearMap.id) : ∃ (g : N →ₗ[R] M), f.comp g = LinearMap.id := by show LinearMap.id ∈ LinearMap.range (LinearMap.llcomp R N M N f) - have inst₁ (I : Ideal R) [I.IsMaximal] : - IsLocalizedModule I.primeCompl (LocalizedModule.map (M := N) (N := N) I.primeCompl) := - inferInstance - have inst₂ (I : Ideal R) [I.IsMaximal] : - IsLocalizedModule I.primeCompl (LocalizedModule.map (M := N) (N := M) I.primeCompl) := - inferInstance - apply - @Submodule.mem_of_localization_maximal R (N →ₗ[R] N) _ _ _ - (fun P _ ↦ Localization.AtPrime P) _ _ _ _ _ _ _ _ - (fun P _ ↦ LocalizedModule.map P.primeCompl) - (fun P _ ↦ inst₁ P) - intro I hI + refine Submodule.mem_of_localization_maximal _ (fun P _ ↦ LocalizedModule.map P.primeCompl) _ _ + fun I hI ↦ ?_ rw [LocalizedModule.map_id] have : LinearMap.id ∈ LinearMap.range (LinearMap.llcomp _ (LocalizedModule I.primeCompl N) _ _ (LocalizedModule.map I.primeCompl f)) := H I hI diff --git a/Mathlib/RingTheory/LocalProperties/Submodule.lean b/Mathlib/RingTheory/LocalProperties/Submodule.lean index f4e9274c697bd..2ecaa20be887a 100644 --- a/Mathlib/RingTheory/LocalProperties/Submodule.lean +++ b/Mathlib/RingTheory/LocalProperties/Submodule.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang, David Swinarski -/ import Mathlib.Algebra.Module.LocalizedModule.Submodule -import Mathlib.RingTheory.Ideal.Colon import Mathlib.RingTheory.Localization.AtPrime /-! @@ -15,87 +14,113 @@ In this file, we show that several conditions on submodules can be checked on st open scoped nonZeroDivisors -variable {R A M} [CommRing R] [CommRing A] [AddCommGroup M] [Algebra R A] [Module R M] [Module A M] +variable {R M M₁ : Type*} [CommSemiring R] [AddCommMonoid M] [Module R M] + [AddCommMonoid M₁] [Module R M₁] + +section maximal variable (Rₚ : ∀ (P : Ideal R) [P.IsMaximal], Type*) - [∀ (P : Ideal R) [P.IsMaximal], CommRing (Rₚ P)] + [∀ (P : Ideal R) [P.IsMaximal], CommSemiring (Rₚ P)] [∀ (P : Ideal R) [P.IsMaximal], Algebra R (Rₚ P)] [∀ (P : Ideal R) [P.IsMaximal], IsLocalization.AtPrime (Rₚ P) P] (Mₚ : ∀ (P : Ideal R) [P.IsMaximal], Type*) - [∀ (P : Ideal R) [P.IsMaximal], AddCommGroup (Mₚ P)] + [∀ (P : Ideal R) [P.IsMaximal], AddCommMonoid (Mₚ P)] [∀ (P : Ideal R) [P.IsMaximal], Module R (Mₚ P)] [∀ (P : Ideal R) [P.IsMaximal], Module (Rₚ P) (Mₚ P)] [∀ (P : Ideal R) [P.IsMaximal], IsScalarTower R (Rₚ P) (Mₚ P)] (f : ∀ (P : Ideal R) [P.IsMaximal], M →ₗ[R] Mₚ P) - [inst : ∀ (P : Ideal R) [P.IsMaximal], IsLocalizedModule P.primeCompl (f P)] + [∀ (P : Ideal R) [P.IsMaximal], IsLocalizedModule P.primeCompl (f P)] + (M₁ₚ : ∀ (P : Ideal R) [P.IsMaximal], Type*) + [∀ (P : Ideal R) [P.IsMaximal], AddCommMonoid (M₁ₚ P)] + [∀ (P : Ideal R) [P.IsMaximal], Module R (M₁ₚ P)] + (f₁ : ∀ (P : Ideal R) [P.IsMaximal], M₁ →ₗ[R] M₁ₚ P) + [∀ (P : Ideal R) [P.IsMaximal], IsLocalizedModule P.primeCompl (f₁ P)] + +theorem Submodule.mem_of_localization_maximal (m : M) (N : Submodule R M) + (h : ∀ (P : Ideal R) [P.IsMaximal], f P m ∈ N.localized₀ P.primeCompl (f P)) : + m ∈ N := by + let I : Ideal R := N.comap (LinearMap.toSpanSingleton R M m) + suffices I = ⊤ by simpa [I] using I.eq_top_iff_one.mp this + refine Not.imp_symm I.exists_le_maximal fun ⟨P, hP, le⟩ ↦ ?_ + obtain ⟨a, ha, s, e⟩ := h P + rw [← IsLocalizedModule.mk'_one P.primeCompl, IsLocalizedModule.mk'_eq_mk'_iff] at e + obtain ⟨t, ht⟩ := e + simp_rw [smul_smul] at ht + exact (t * s).2 (le <| by apply ht ▸ smul_mem _ _ ha) /-- Let `N₁ N₂ : Submodule R M`. If the localization of `N₁` at each maximal ideal `P` is included in the localization of `N₂` at `P`, then `N₁ ≤ N₂`. -/ theorem Submodule.le_of_localization_maximal {N₁ N₂ : Submodule R M} (h : ∀ (P : Ideal R) [P.IsMaximal], - N₁.localized' (Rₚ P) P.primeCompl (f P) ≤ N₂.localized' (Rₚ P) P.primeCompl (f P)) : - N₁ ≤ N₂ := by - intro x hx - suffices N₂.colon (Submodule.span R {x}) = ⊤ by - simpa using Submodule.mem_colon.mp - (show (1 : R) ∈ N₂.colon (Submodule.span R {x}) from this.symm ▸ Submodule.mem_top) x - (Submodule.mem_span_singleton_self x) - refine Not.imp_symm (N₂.colon (Submodule.span R {x})).exists_le_maximal ?_ - push_neg - intro P hP le - obtain ⟨a, ha, s, e⟩ := h P ⟨x, hx, 1, rfl⟩ - rw [IsLocalizedModule.mk'_eq_mk'_iff] at e - obtain ⟨t, ht⟩ := e - simp at ht - refine (t * s).2 (le (Submodule.mem_colon_singleton.mpr ?_)) - simp only [Submonoid.coe_mul, mul_smul, ← Submonoid.smul_def, ht] - exact N₂.smul_mem _ ha + N₁.localized₀ P.primeCompl (f P) ≤ N₂.localized₀ P.primeCompl (f P)) : + N₁ ≤ N₂ := + fun m hm ↦ mem_of_localization_maximal _ f _ _ fun P hP ↦ h P ⟨m, hm, 1, by simp⟩ /-- Let `N₁ N₂ : Submodule R M`. If the localization of `N₁` at each maximal ideal `P` is equal to the localization of `N₂` at `P`, then `N₁ = N₂`. -/ -theorem Submodule.eq_of_localization_maximal {N₁ N₂ : Submodule R M} +theorem Submodule.eq_of_localization₀_maximal {N₁ N₂ : Submodule R M} (h : ∀ (P : Ideal R) [P.IsMaximal], - N₁.localized' (Rₚ P) P.primeCompl (f P) = N₂.localized' (Rₚ P) P.primeCompl (f P)) : + N₁.localized₀ P.primeCompl (f P) = N₂.localized₀ P.primeCompl (f P)) : N₁ = N₂ := - le_antisymm (Submodule.le_of_localization_maximal Rₚ Mₚ f fun P _ => (h P).le) - (Submodule.le_of_localization_maximal Rₚ Mₚ f fun P _ => (h P).ge) + le_antisymm (Submodule.le_of_localization_maximal Mₚ f fun P _ ↦ (h P).le) + (Submodule.le_of_localization_maximal Mₚ f fun P _ ↦ (h P).ge) /-- A submodule is trivial if its localization at every maximal ideal is trivial. -/ -theorem Submodule.eq_bot_of_localization_maximal (N₁ : Submodule R M) +theorem Submodule.eq_bot_of_localization₀_maximal (N : Submodule R M) + (h : ∀ (P : Ideal R) [P.IsMaximal], N.localized₀ P.primeCompl (f P) = ⊥) : + N = ⊥ := + Submodule.eq_of_localization₀_maximal Mₚ f fun P hP ↦ by simpa using h P + +theorem Submodule.eq_top_of_localization₀_maximal (N : Submodule R M) + (h : ∀ (P : Ideal R) [P.IsMaximal], N.localized₀ P.primeCompl (f P) = ⊤) : + N = ⊤ := + Submodule.eq_of_localization₀_maximal Mₚ f fun P hP ↦ by simpa using h P + +theorem Module.eq_of_localization_maximal (m m' : M) + (h : ∀ (P : Ideal R) [P.IsMaximal], f P m = f P m') : + m = m' := by + by_contra! ne + rw [← one_smul R m, ← one_smul R m'] at ne + have ⟨P, mP, le⟩ := (eqIdeal R m m').exists_le_maximal ((Ideal.ne_top_iff_one _).mpr ne) + have ⟨s, hs⟩ := (IsLocalizedModule.eq_iff_exists P.primeCompl _).mp (h P) + exact s.2 (le hs) + +theorem Module.eq_zero_of_localization_maximal (m : M) + (h : ∀ (P : Ideal R) [P.IsMaximal], f P m = 0) : + m = 0 := + eq_of_localization_maximal _ f _ _ fun P _ ↦ by rw [h, map_zero] + +theorem LinearMap.eq_of_localization_maximal (g g' : M →ₗ[R] M₁) (h : ∀ (P : Ideal R) [P.IsMaximal], - N₁.localized' (Rₚ P) P.primeCompl (f P) = ⊥) : - N₁ = ⊥ := - Submodule.eq_of_localization_maximal Rₚ Mₚ f fun P hP => by simpa using h P - -theorem Submodule.mem_of_localization_maximal (r : M) (N₁ : Submodule R M) - (h : ∀ (P : Ideal R) [P.IsMaximal], f P r ∈ N₁.localized' (Rₚ P) P.primeCompl (f P)) : - r ∈ N₁ := by - rw [← SetLike.mem_coe, ← Set.singleton_subset_iff, ← Submodule.span_le] - apply Submodule.le_of_localization_maximal Rₚ Mₚ f - intro N₂ hJ - rw [Submodule.localized'_span, Submodule.span_le, Set.image_subset_iff, Set.singleton_subset_iff] - exact h N₂ - -include Rₚ in -theorem Module.eq_zero_of_localization_maximal (r : M) - (h : ∀ (P : Ideal R) [P.IsMaximal], f P r = 0) : - r = 0 := by - rw [← Submodule.mem_bot (R := R)] - apply Submodule.mem_of_localization_maximal Rₚ Mₚ f r ⊥ (by simpa using h) - -include Rₚ in -theorem Module.eq_of_localization_maximal (r s : M) - (h : ∀ (P : Ideal R) [P.IsMaximal], f P r = f P s) : - r = s := by - rw [← sub_eq_zero] - simp_rw [← @sub_eq_zero _ _ (f _ _), ← map_sub] at h - exact Module.eq_zero_of_localization_maximal Rₚ Mₚ f _ h - -include Rₚ f in + IsLocalizedModule.map P.primeCompl (f P) (f₁ P) g = + IsLocalizedModule.map P.primeCompl (f P) (f₁ P) g') : + g = g' := + ext fun x ↦ Module.eq_of_localization_maximal _ f₁ _ _ fun P _ ↦ by + simpa only [IsLocalizedModule.map_apply] using DFunLike.congr_fun (h P) (f P x) + +include f in theorem Module.subsingleton_of_localization_maximal (h : ∀ (P : Ideal R) [P.IsMaximal], Subsingleton (Mₚ P)) : Subsingleton M := by rw [subsingleton_iff_forall_eq 0] intro x - exact Module.eq_of_localization_maximal Rₚ Mₚ f x 0 fun _ _ ↦ Subsingleton.elim _ _ + exact Module.eq_of_localization_maximal Mₚ f x 0 fun _ _ ↦ Subsingleton.elim _ _ + +theorem Submodule.eq_of_localization_maximal {N₁ N₂ : Submodule R M} + (h : ∀ (P : Ideal R) [P.IsMaximal], + N₁.localized' (Rₚ P) P.primeCompl (f P) = N₂.localized' (Rₚ P) P.primeCompl (f P)) : + N₁ = N₂ := + eq_of_localization₀_maximal Mₚ f fun P _ ↦ congr(restrictScalars _ $(h P)) + +theorem Submodule.eq_bot_of_localization_maximal (N : Submodule R M) + (h : ∀ (P : Ideal R) [P.IsMaximal], N.localized' (Rₚ P) P.primeCompl (f P) = ⊥) : + N = ⊥ := + Submodule.eq_of_localization_maximal Rₚ Mₚ f fun P hP ↦ by simpa using h P + +theorem Submodule.eq_top_of_localization_maximal (N : Submodule R M) + (h : ∀ (P : Ideal R) [P.IsMaximal], N.localized' (Rₚ P) P.primeCompl (f P) = ⊤) : + N = ⊤ := + Submodule.eq_of_localization_maximal Rₚ Mₚ f fun P hP ↦ by simpa using h P + +end maximal From 8e0d70c35dcdbe42cd6fa06c579a118735380526 Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 20 Nov 2024 18:07:47 +0000 Subject: [PATCH 159/829] fix: `>>` and print author (#19300) Use `>>` to set `GITHUB_OUTPUT`, instead of `>`. --- .github/workflows/maintainer_bors.yml | 7 ++++--- 1 file changed, 4 insertions(+), 3 deletions(-) diff --git a/.github/workflows/maintainer_bors.yml b/.github/workflows/maintainer_bors.yml index d81dded626572..6832b6e007eed 100644 --- a/.github/workflows/maintainer_bors.yml +++ b/.github/workflows/maintainer_bors.yml @@ -42,16 +42,17 @@ jobs: sed -n 's=^bors *\(merge\|r+\) *$=ready-to-merge=p; s=^bors *d.*=delegated=p' | head -1)" printf $'"bors delegate" or "bors merge" found? \'%s\'\n' "${m_or_d}" + printf $'AUTHOR: \'%s\'\n' "${AUTHOR}" - printf $'mOrD=%s\n' "${m_or_d}" > "${GITHUB_OUTPUT}" + printf $'mOrD=%s\n' "${m_or_d}" >> "${GITHUB_OUTPUT}" if [ "${AUTHOR}" == 'leanprover-community-mathlib4-bot' ] || [ "${AUTHOR}" == 'leanprover-community-bot-assistant' ] then printf $'bot=true\n' - printf $'bot=true\n' > "${GITHUB_OUTPUT}" + printf $'bot=true\n' >> "${GITHUB_OUTPUT}" else printf $'bot=false\n' - printf $'bot=false\n' > "${GITHUB_OUTPUT}" + printf $'bot=false\n' >> "${GITHUB_OUTPUT}" fi - name: Check whether user is a mathlib admin From 9f24be2a7684317446564b4f8656fe70b308c425 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Wed, 20 Nov 2024 19:44:13 +0000 Subject: [PATCH 160/829] feat(AlgebraicGeometry): Proj is separated (#19290) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Patience Ablett Co-authored-by: Kevin Buzzard Co-authored-by: Harald Carlens Co-authored-by: Wayne Ng Kwing King Co-authored-by: Michael Schlößer Co-authored-by: Justus Springer Co-authored-by: Jujian Zhang This contribution was created as part of the Durham Computational Algebraic Geometry Workshop Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib.lean | 1 + .../ProjectiveSpectrum/Basic.lean | 84 ++++++++++++ .../ProjectiveSpectrum/Proper.lean | 126 ++++++++++++++++++ .../ProjectiveSpectrum/Scheme.lean | 10 ++ .../HomogeneousLocalization.lean | 90 +++++++++++++ .../RingTheory/Localization/Away/Basic.lean | 9 ++ 6 files changed, 320 insertions(+) create mode 100644 Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Proper.lean diff --git a/Mathlib.lean b/Mathlib.lean index ec30d813df6af..14b5b0c4d7f7a 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -966,6 +966,7 @@ import Mathlib.AlgebraicGeometry.PrimeSpectrum.Module import Mathlib.AlgebraicGeometry.PrimeSpectrum.Noetherian import Mathlib.AlgebraicGeometry.PrimeSpectrum.TensorProduct import Mathlib.AlgebraicGeometry.ProjectiveSpectrum.Basic +import Mathlib.AlgebraicGeometry.ProjectiveSpectrum.Proper import Mathlib.AlgebraicGeometry.ProjectiveSpectrum.Scheme import Mathlib.AlgebraicGeometry.ProjectiveSpectrum.StructureSheaf import Mathlib.AlgebraicGeometry.ProjectiveSpectrum.Topology diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean index b29a137f71eb2..1ebe7c785f7bb 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean @@ -166,6 +166,90 @@ lemma awayι_toSpecZero : awayι 𝒜 f f_deg hm ≫ toSpecZero 𝒜 = ← Spec.map_comp, ← Spec.map_comp, ← Spec.map_comp] rfl +variable {f} +variable {m' : ℕ} {g : A} (g_deg : g ∈ 𝒜 m') (hm' : 0 < m') {x : A} (hx : x = f * g) + +@[reassoc] +lemma awayMap_awayToSection : + CommRingCat.ofHom (awayMap 𝒜 g_deg hx) ≫ awayToSection 𝒜 x = + awayToSection 𝒜 f ≫ (Proj 𝒜).presheaf.map (homOfLE (basicOpen_mono _ _ _ ⟨_, hx⟩)).op := by + ext a + apply Subtype.ext + ext ⟨i, hi⟩ + obtain ⟨⟨n, a, ⟨b, hb'⟩, i, rfl : _ = b⟩, rfl⟩ := mk_surjective a + simp only [CommRingCat.forget_obj, CommRingCat.coe_of, CommRingCat.ofHom, CommRingCat.coe_comp_of, + RingHom.coe_comp, Function.comp_apply, homOfLE_leOfHom] + erw [ProjectiveSpectrum.Proj.awayToSection_apply] + rw [val_awayMap_mk, Localization.mk_eq_mk', IsLocalization.map_mk', + ← Localization.mk_eq_mk'] + refine Localization.mk_eq_mk_iff.mpr ?_ + rw [Localization.r_iff_exists] + use 1 + simp only [OneMemClass.coe_one, RingHom.id_apply, one_mul, hx] + ring + +@[reassoc] +lemma basicOpenToSpec_SpecMap_awayMap : + basicOpenToSpec 𝒜 x ≫ Spec.map (CommRingCat.ofHom (awayMap 𝒜 g_deg hx)) = + (Proj 𝒜).homOfLE (basicOpen_mono _ _ _ ⟨_, hx⟩) ≫ basicOpenToSpec 𝒜 f := by + rw [basicOpenToSpec, Category.assoc, ← Spec.map_comp, awayMap_awayToSection, + Spec.map_comp, Scheme.Opens.toSpecΓ_SpecMap_map_assoc] + rfl + +@[reassoc] +lemma SpecMap_awayMap_awayι : + Spec.map (CommRingCat.ofHom (awayMap 𝒜 g_deg hx)) ≫ awayι 𝒜 f f_deg hm = + awayι 𝒜 x (hx ▸ SetLike.mul_mem_graded f_deg g_deg) (hm.trans_le (m.le_add_right m')) := by + rw [awayι, awayι, Iso.eq_inv_comp, basicOpenIsoSpec_hom, basicOpenToSpec_SpecMap_awayMap_assoc, + ← basicOpenIsoSpec_hom _ _ f_deg hm, Iso.hom_inv_id_assoc, Scheme.homOfLE_ι] + +/-- The isomorphism `D₊(f) ×[Proj 𝒜] D₊(g) ≅ D₊(fg)`. -/ +noncomputable +def pullbackAwayιIso : + Limits.pullback (awayι 𝒜 f f_deg hm) (awayι 𝒜 g g_deg hm') ≅ + Spec (CommRingCat.of (Away 𝒜 x)) := + IsOpenImmersion.isoOfRangeEq (Limits.pullback.fst _ _ ≫ awayι 𝒜 f f_deg hm) + (awayι 𝒜 x (hx ▸ SetLike.mul_mem_graded f_deg g_deg) (hm.trans_le (m.le_add_right m'))) <| by + rw [IsOpenImmersion.range_pullback_to_base_of_left] + show ((awayι 𝒜 f _ _).opensRange ⊓ (awayι 𝒜 g _ _).opensRange).1 = (awayι 𝒜 _ _ _).opensRange.1 + rw [opensRange_awayι, opensRange_awayι, opensRange_awayι, ← basicOpen_mul, hx] + +@[reassoc (attr := simp)] +lemma pullbackAwayιIso_hom_awayι : + (pullbackAwayιIso 𝒜 f_deg hm g_deg hm' hx).hom ≫ + awayι 𝒜 x (hx ▸ SetLike.mul_mem_graded f_deg g_deg) (hm.trans_le (m.le_add_right m')) = + Limits.pullback.fst _ _ ≫ awayι 𝒜 f f_deg hm := + IsOpenImmersion.isoOfRangeEq_hom_fac .. + +@[reassoc (attr := simp)] +lemma pullbackAwayιIso_hom_SpecMap_awayMap_left : + (pullbackAwayιIso 𝒜 f_deg hm g_deg hm' hx).hom ≫ + Spec.map (CommRingCat.ofHom (awayMap 𝒜 g_deg hx)) = Limits.pullback.fst _ _ := by + rw [← cancel_mono (awayι 𝒜 f f_deg hm), ← pullbackAwayιIso_hom_awayι, + Category.assoc, SpecMap_awayMap_awayι] + +@[reassoc (attr := simp)] +lemma pullbackAwayιIso_hom_SpecMap_awayMap_right : + (pullbackAwayιIso 𝒜 f_deg hm g_deg hm' hx).hom ≫ + Spec.map (CommRingCat.ofHom (awayMap 𝒜 f_deg (hx.trans (mul_comm _ _)))) = + Limits.pullback.snd _ _ := by + rw [← cancel_mono (awayι 𝒜 g g_deg hm'), ← Limits.pullback.condition, + ← pullbackAwayιIso_hom_awayι 𝒜 f_deg hm g_deg hm' hx, + Category.assoc, SpecMap_awayMap_awayι] + rfl + +@[reassoc (attr := simp)] +lemma pullbackAwayιIso_inv_fst : + (pullbackAwayιIso 𝒜 f_deg hm g_deg hm' hx).inv ≫ Limits.pullback.fst _ _ = + Spec.map (CommRingCat.ofHom (awayMap 𝒜 g_deg hx)) := by + rw [← pullbackAwayιIso_hom_SpecMap_awayMap_left, Iso.inv_hom_id_assoc] + +@[reassoc (attr := simp)] +lemma pullbackAwayιIso_inv_snd : + (pullbackAwayιIso 𝒜 f_deg hm g_deg hm' hx).inv ≫ Limits.pullback.snd _ _ = + Spec.map (CommRingCat.ofHom (awayMap 𝒜 f_deg (hx.trans (mul_comm _ _)))) := by + rw [← pullbackAwayιIso_hom_SpecMap_awayMap_right, Iso.inv_hom_id_assoc] + open TopologicalSpace.Opens in /-- Given a family of homogeneous elements `f` of positive degree that spans the irrelavent ideal, `Spec (A_f)₀ ⟶ Proj A` forms an affine open cover of `Proj A`. -/ diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Proper.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Proper.lean new file mode 100644 index 0000000000000..6235852b32c97 --- /dev/null +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Proper.lean @@ -0,0 +1,126 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Patience Ablett, Kevin Buzzard, Harald Carlens, Wayne Ng Kwing King, Michael Schlößer, + Justus Springer, Andrew Yang, Jujian Zhang +-/ +import Mathlib.AlgebraicGeometry.Morphisms.Separated +import Mathlib.AlgebraicGeometry.ProjectiveSpectrum.Basic + +/-! +# Properness of `Proj A` + +We show that `Proj 𝒜` is a separated scheme. + +## TODO +- Show that `Proj 𝒜` satisfies the valuative criterion. + +## Notes +This contribution was created as part of the Durham Computational Algebraic Geometry Workshop + +-/ + +namespace AlgebraicGeometry.Proj + +variable {R A : Type*} +variable [CommRing R] [CommRing A] [Algebra R A] +variable (𝒜 : ℕ → Submodule R A) +variable [GradedAlgebra 𝒜] + +open Scheme CategoryTheory Limits pullback HomogeneousLocalization + +lemma lift_awayMapₐ_awayMapₐ_surjective {d e : ℕ} {f : A} (hf : f ∈ 𝒜 d) + {g : A} (hg : g ∈ 𝒜 e) {x : A} (hx : x = f * g) (hd : 0 < d) : + Function.Surjective + (Algebra.TensorProduct.lift (awayMapₐ 𝒜 hg hx) (awayMapₐ 𝒜 hf (hx.trans (mul_comm _ _))) + (fun _ _ ↦ .all _ _)) := by + intro z + obtain ⟨⟨n, ⟨a, ha⟩, ⟨b, hb'⟩, ⟨j, (rfl : _ = b)⟩⟩, rfl⟩ := mk_surjective z + by_cases hfg : (f * g) ^ j = 0 + · use 0 + have := HomogeneousLocalization.subsingleton 𝒜 (x := Submonoid.powers x) ⟨j, hx ▸ hfg⟩ + exact this.elim _ _ + have : n = j * (d + e) := by + apply DirectSum.degree_eq_of_mem_mem 𝒜 hb' + convert SetLike.pow_mem_graded _ _ using 2 + · infer_instance + · exact hx ▸ SetLike.mul_mem_graded hf hg + · exact hx ▸ hfg + let x0 : NumDenSameDeg 𝒜 (.powers f) := + { deg := j * (d * (e + 1)) + num := ⟨a * g ^ (j * (d - 1)), by + convert SetLike.mul_mem_graded ha (SetLike.pow_mem_graded _ hg) using 2 + rw [this] + cases d + · contradiction + · simp; ring⟩ + den := ⟨f ^ (j * (e + 1)), by convert SetLike.pow_mem_graded _ hf using 2; ring⟩ + den_mem := ⟨_,rfl⟩ } + let y0 : NumDenSameDeg 𝒜 (.powers g) := + { deg := j * (d * e) + num := ⟨f ^ (j * e), by convert SetLike.pow_mem_graded _ hf using 2; ring⟩ + den := ⟨g ^ (j * d), by convert SetLike.pow_mem_graded _ hg using 2; ring⟩ + den_mem := ⟨_,rfl⟩ } + use mk x0 ⊗ₜ mk y0 + ext + simp only [Algebra.TensorProduct.lift_tmul, awayMapₐ_apply, val_mul, + val_awayMap_mk, Localization.mk_mul, val_mk, x0, y0] + rw [Localization.mk_eq_mk_iff, Localization.r_iff_exists] + use 1 + simp only [OneMemClass.coe_one, one_mul, Submonoid.mk_mul_mk] + cases d + · contradiction + · simp only [hx, add_tsub_cancel_right] + ring + +open TensorProduct in +instance isSeparated : IsSeparated (toSpecZero 𝒜) := by + refine ⟨IsLocalAtTarget.of_openCover (Pullback.openCoverOfLeftRight + (affineOpenCover 𝒜).openCover (affineOpenCover 𝒜).openCover _ _) ?_⟩ + intro ⟨i, j⟩ + dsimp [Scheme, Cover.pullbackHom] + refine (MorphismProperty.cancel_left_of_respectsIso (P := @IsClosedImmersion) + (f := (pullbackDiagonalMapIdIso ..).inv) _).mp ?_ + let e₁ : pullback ((affineOpenCover 𝒜).map i ≫ toSpecZero 𝒜) + ((affineOpenCover 𝒜).map j ≫ toSpecZero 𝒜) ≅ + Spec (.of <| TensorProduct (𝒜 0) (Away 𝒜 i.2) (Away 𝒜 j.2)) := by + refine pullback.congrHom ?_ ?_ ≪≫ pullbackSpecIso (𝒜 0) (Away 𝒜 i.2) (Away 𝒜 j.2) + · simp [affineOpenCover, openCoverOfISupEqTop, awayι_toSpecZero]; rfl + · simp [affineOpenCover, openCoverOfISupEqTop, awayι_toSpecZero]; rfl + let e₂ : pullback ((affineOpenCover 𝒜).map i) ((affineOpenCover 𝒜).map j) ≅ + Spec (.of <| (Away 𝒜 (i.2 * j.2))) := + pullbackAwayιIso 𝒜 _ _ _ _ rfl + rw [← MorphismProperty.cancel_right_of_respectsIso (P := @IsClosedImmersion) _ e₁.hom, + ← MorphismProperty.cancel_left_of_respectsIso (P := @IsClosedImmersion) e₂.inv] + let F : Away 𝒜 i.2.1 ⊗[𝒜 0] Away 𝒜 j.2.1 →+* Away 𝒜 (i.2.1 * j.2.1) := + (Algebra.TensorProduct.lift (awayMapₐ 𝒜 j.2.2 rfl) (awayMapₐ 𝒜 i.2.2 (mul_comm _ _)) + (fun _ _ ↦ .all _ _)).toRingHom + have : Function.Surjective F := lift_awayMapₐ_awayMapₐ_surjective 𝒜 i.2.2 j.2.2 rfl i.1.2 + convert IsClosedImmersion.spec_of_surjective + (CommRingCat.ofHom (R := Away 𝒜 i.2.1 ⊗[𝒜 0] Away 𝒜 j.2.1) F) this using 1 + rw [← cancel_mono (pullbackSpecIso ..).inv] + apply pullback.hom_ext + · simp only [Iso.trans_hom, congrHom_hom, Category.assoc, Iso.hom_inv_id, Category.comp_id, + limit.lift_π, id_eq, eq_mpr_eq_cast, PullbackCone.mk_pt, PullbackCone.mk_π_app, e₂, e₁, + pullbackDiagonalMapIdIso_inv_snd_fst, AlgHom.toRingHom_eq_coe, pullbackSpecIso_inv_fst, + ← Spec.map_comp] + erw [pullbackAwayιIso_inv_fst] + congr 1 + ext x + exact DFunLike.congr_fun (Algebra.TensorProduct.lift_comp_includeLeft + (awayMapₐ 𝒜 j.2.2 rfl) (awayMapₐ 𝒜 i.2.2 (mul_comm _ _)) (fun _ _ ↦ .all _ _)).symm x + · simp only [Iso.trans_hom, congrHom_hom, Category.assoc, Iso.hom_inv_id, Category.comp_id, + limit.lift_π, id_eq, eq_mpr_eq_cast, PullbackCone.mk_pt, PullbackCone.mk_π_app, + pullbackDiagonalMapIdIso_inv_snd_snd, AlgHom.toRingHom_eq_coe, pullbackSpecIso_inv_snd, ← + Spec.map_comp, e₂, e₁] + erw [pullbackAwayιIso_inv_snd] + congr 1 + ext x + exact DFunLike.congr_fun (Algebra.TensorProduct.lift_comp_includeRight + (awayMapₐ 𝒜 j.2.2 rfl) (awayMapₐ 𝒜 i.2.2 (mul_comm _ _)) (fun _ _ ↦ .all _ _)).symm x + +@[stacks 01MC] +instance : Scheme.IsSeparated (Proj 𝒜) := + (HasAffineProperty.iff_of_isAffine (P := @IsSeparated)).mp (isSeparated 𝒜) + +end AlgebraicGeometry.Proj diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean index 462689531697b..781f0783739f0 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean @@ -603,6 +603,16 @@ lemma awayToSection_germ (f x hx) : apply (Proj.stalkIso' 𝒜 x).eq_symm_apply.mpr apply Proj.stalkIso'_germ +lemma awayToSection_apply (f : A) (x p) : + (((ProjectiveSpectrum.Proj.awayToSection 𝒜 f).1 x).val p).val = + IsLocalization.map (M := Submonoid.powers f) (T := p.1.1.toIdeal.primeCompl) _ + (RingHom.id _) (Submonoid.powers_le.mpr p.2) x.val := by + obtain ⟨x, rfl⟩ := HomogeneousLocalization.mk_surjective x + show (HomogeneousLocalization.mapId 𝒜 _ _).val = _ + dsimp [HomogeneousLocalization.mapId, HomogeneousLocalization.map] + rw [Localization.mk_eq_mk', Localization.mk_eq_mk', IsLocalization.map_mk'] + rfl + /-- The ring map from `A⁰_ f` to the global sections of the structure sheaf of the projective spectrum of `A` restricted to the basic open set `D(f)`. diff --git a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean index dcab8f6f958a6..af6dc235b4928 100644 --- a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean +++ b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean @@ -6,6 +6,7 @@ Authors: Jujian Zhang, Eric Wieser import Mathlib.Order.Filter.AtTopBot import Mathlib.RingTheory.Localization.AtPrime import Mathlib.RingTheory.GradedAlgebra.Basic +import Mathlib.RingTheory.Localization.Away.Basic /-! # Homogeneous Localization @@ -473,6 +474,11 @@ def fromZeroRingHom : 𝒜 0 →+* HomogeneousLocalization 𝒜 x where map_zero' := rfl map_add' f g := by ext; simp [Localization.add_mk, add_comm f.1 g.1] +instance : Algebra (𝒜 0) (HomogeneousLocalization 𝒜 x) := + (fromZeroRingHom 𝒜 x).toAlgebra + +lemma algebraMap_eq : algebraMap (𝒜 0) (HomogeneousLocalization 𝒜 x) = fromZeroRingHom 𝒜 x := rfl + end HomogeneousLocalization namespace HomogeneousLocalization @@ -627,4 +633,88 @@ lemma map_mk (g : A →+* B) end +section mapAway + +variable [AddCommMonoid ι] [DecidableEq ι] [GradedAlgebra 𝒜] +variable {d e : ι} {f : A} (hf : f ∈ 𝒜 d) {g : A} (hg : g ∈ 𝒜 e) +variable {x : A} (hx : x = f * g) + +variable (𝒜) + +/-- Given `f ∣ x`, this is the map `A_{(f)} → A_f → A_x`. We will lift this to a map +`A_{(f)} → A_{(x)}` in `awayMap`. -/ +private def awayMapAux (hx : f ∣ x) : Away 𝒜 f →+* Localization.Away x := + (Localization.awayLift (algebraMap A _) _ + (isUnit_of_dvd_unit (map_dvd _ hx) (IsLocalization.Away.algebraMap_isUnit x))).comp + (algebraMap (Away 𝒜 f) (Localization.Away f)) + +lemma awayMapAux_mk (n a i hi) : + awayMapAux 𝒜 ⟨_, hx⟩ (mk ⟨n, a, ⟨f ^ i, hi⟩, ⟨i, rfl⟩⟩) = + Localization.mk (a * g ^ i) ⟨x ^ i, (Submonoid.mem_powers_iff _ _).mpr ⟨i, rfl⟩⟩ := by + have : algebraMap A (Localization.Away x) f * + (Localization.mk g ⟨f * g, (Submonoid.mem_powers_iff _ _).mpr ⟨1, by simp [hx]⟩⟩) = 1 := by + rw [← Algebra.smul_def, Localization.smul_mk] + exact Localization.mk_self ⟨f*g, _⟩ + simp [awayMapAux] + rw [Localization.awayLift_mk (hv := this), ← Algebra.smul_def, + Localization.mk_pow, Localization.smul_mk] + subst hx + rfl + +include hg in +lemma range_awayMapAux_subset : + Set.range (awayMapAux 𝒜 (f := f) ⟨_, hx⟩) ⊆ Set.range (val (𝒜 := 𝒜)) := by + rintro _ ⟨z, rfl⟩ + obtain ⟨⟨n, ⟨a, ha⟩, ⟨b, hb'⟩, j, rfl : _ = b⟩, rfl⟩ := mk_surjective z + use mk ⟨n+j•e,⟨a*g^j, ?_⟩ ,⟨x^j, ?_⟩, j, rfl⟩ + · simp [awayMapAux_mk 𝒜 (hx := hx)] + · apply SetLike.mul_mem_graded ha + exact SetLike.pow_mem_graded _ hg + · rw [hx, mul_pow] + apply SetLike.mul_mem_graded hb' + exact SetLike.pow_mem_graded _ hg + +/-- Given `x = f * g` with `g` homogeneous of positive degree, +this is the map `A_{(f)} → A_{(x)}` taking `a/f^i` to `ag^i/(fg)^i`. -/ +def awayMap : Away 𝒜 f →+* Away 𝒜 x := by + let e := RingEquiv.ofLeftInverse (f := algebraMap (Away 𝒜 x) (Localization.Away x)) + (h := (val_injective _).hasLeftInverse.choose_spec) + refine RingHom.comp (e.symm.toRingHom.comp (Subring.inclusion ?_)) + (awayMapAux 𝒜 (f := f) ⟨_, hx⟩).rangeRestrict + exact range_awayMapAux_subset 𝒜 hg hx + +lemma val_awayMap_eq_aux (a) : (awayMap 𝒜 hg hx a).val = awayMapAux 𝒜 ⟨_, hx⟩ a := by + let e := RingEquiv.ofLeftInverse (f := algebraMap (Away 𝒜 x) (Localization.Away x)) + (h := (val_injective _).hasLeftInverse.choose_spec) + dsimp [awayMap] + convert_to (e (e.symm ⟨awayMapAux 𝒜 (f := f) ⟨_, hx⟩ a, + range_awayMapAux_subset 𝒜 hg hx ⟨_, rfl⟩⟩)).1 = _ + rw [e.apply_symm_apply] + +lemma val_awayMap (a) : (awayMap 𝒜 hg hx a).val = Localization.awayLift (algebraMap A _) _ + (isUnit_of_dvd_unit (map_dvd _ ⟨_, hx⟩) (IsLocalization.Away.algebraMap_isUnit x)) a.val := by + rw [val_awayMap_eq_aux] + rfl + +lemma awayMap_fromZeroRingHom (a) : + awayMap 𝒜 hg hx (fromZeroRingHom 𝒜 _ a) = fromZeroRingHom 𝒜 _ a := by + ext + simp only [fromZeroRingHom, RingHom.coe_mk, MonoidHom.coe_mk, OneHom.coe_mk, + val_awayMap, val_mk, SetLike.GradeZero.coe_one] + convert IsLocalization.lift_eq _ _ + +lemma val_awayMap_mk (n a i hi) : (awayMap 𝒜 hg hx (mk ⟨n, a, ⟨f ^ i, hi⟩, ⟨i, rfl⟩⟩)).val = + Localization.mk (a * g ^ i) ⟨x ^ i, (Submonoid.mem_powers_iff _ _).mpr ⟨i, rfl⟩⟩ := by + rw [val_awayMap_eq_aux, awayMapAux_mk 𝒜 (hx := hx)] + +/-- Given `x = f * g` with `g` homogeneous of positive degree, +this is the map `A_{(f)} → A_{(x)}` taking `a/f^i` to `ag^i/(fg)^i`. -/ +def awayMapₐ : Away 𝒜 f →ₐ[𝒜 0] Away 𝒜 x where + __ := awayMap 𝒜 hg hx + commutes' _ := awayMap_fromZeroRingHom .. + +@[simp] lemma awayMapₐ_apply (a) : awayMapₐ 𝒜 hg hx a = awayMap 𝒜 hg hx a := rfl + +end mapAway + end HomogeneousLocalization diff --git a/Mathlib/RingTheory/Localization/Away/Basic.lean b/Mathlib/RingTheory/Localization/Away/Basic.lean index d587619c72796..01a01a331e08b 100644 --- a/Mathlib/RingTheory/Localization/Away/Basic.lean +++ b/Mathlib/RingTheory/Localization/Away/Basic.lean @@ -361,6 +361,15 @@ noncomputable abbrev awayLift (f : R →+* P) (r : R) (hr : IsUnit (f r)) : Localization.Away r →+* P := IsLocalization.Away.lift r hr +lemma awayLift_mk {A : Type*} [CommRing A] (f : R →+* A) (r : R) + (a : R) (v : A) (hv : f r * v = 1) (j : ℕ) : + Localization.awayLift f r (isUnit_iff_exists_inv.mpr ⟨v, hv⟩) + (Localization.mk a ⟨r ^ j, j, rfl⟩) = f a * v ^ j := by + rw [Localization.mk_eq_mk'] + erw [IsLocalization.lift_mk'] + rw [Units.mul_inv_eq_iff_eq_mul] + simp [IsUnit.liftRight, mul_assoc, ← mul_pow, (mul_comm _ _).trans hv] + /-- Given a map `f : R →+* S` and an element `r : R`, we may construct a map `Rᵣ →+* Sᵣ`. -/ noncomputable abbrev awayMap (f : R →+* P) (r : R) : Localization.Away r →+* Localization.Away (f r) := From bb8094aa6186788cfc45f171d49e69a785cadb5f Mon Sep 17 00:00:00 2001 From: Pietro Monticone <38562595+pitmonticone@users.noreply.github.com> Date: Wed, 20 Nov 2024 20:06:57 +0000 Subject: [PATCH 161/829] feat(Topology/MetricSpace/Pseudo/Defs): add `dense_iff_iUnion_ball` and `dist_eq_of_dist_zero` (#19294) Upstreamed from the [Carleson](https://github.com/fpvandoorn/carleson) project. Co-authored-by: Floris van Doorn --- Mathlib/Topology/MetricSpace/Pseudo/Defs.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean b/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean index 82a80a1ff7f93..6f5392e44495c 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean @@ -1080,6 +1080,10 @@ theorem tendsto_iff_of_dist {f₁ f₂ : ι → α} {p : Filter ι} {a : α} end Real +theorem PseudoMetricSpace.dist_eq_of_dist_zero (x : α) {y z : α} (h : dist y z = 0) : + dist x y = dist x z := + dist_comm y x ▸ dist_comm z x ▸ sub_eq_zero.1 (abs_nonpos_iff.1 (h ▸ abs_dist_sub_le y z x)) + -- Porting note: 3 new lemmas theorem dist_dist_dist_le_left (x y z : α) : dist (dist x z) (dist y z) ≤ dist x y := abs_dist_sub_le .. @@ -1127,6 +1131,10 @@ theorem dense_iff {s : Set α} : Dense s ↔ ∀ x, ∀ r > 0, (ball x r ∩ s). forall_congr' fun x => by simp only [mem_closure_iff, Set.Nonempty, exists_prop, mem_inter_iff, mem_ball', and_comm] +theorem dense_iff_iUnion_ball (s : Set α) : Dense s ↔ ∀ r > 0, ⋃ c ∈ s, ball c r = univ := by + simp_rw [eq_univ_iff_forall, mem_iUnion, exists_prop, mem_ball, Dense, mem_closure_iff, + forall_comm (α := α)] + theorem denseRange_iff {f : β → α} : DenseRange f ↔ ∀ x, ∀ r > 0, ∃ y, dist x (f y) < r := forall_congr' fun x => by simp only [mem_closure_iff, exists_range_iff] From 3c79dac7ee6ac164a9919f3e21ab38d3e05aef82 Mon Sep 17 00:00:00 2001 From: Scott Carnahan <128885296+ScottCarnahan@users.noreply.github.com> Date: Wed, 20 Nov 2024 20:34:20 +0000 Subject: [PATCH 162/829] feat (LinearAlgebra/RootSystem/Finite): nondegeneracy of canonical bilinear form restricted to root span (#18569) The canonical bilinear form for a root pairing over a linearly ordered commutative ring is nondegenerate when restricted to the span of roots. Co-authored-by: leanprover-community-mathlib4-bot --- Mathlib.lean | 1 + .../Order/BigOperators/Ring/Finset.lean | 14 +- Mathlib/LinearAlgebra/Dimension/Basic.lean | 7 + .../LinearAlgebra/Dimension/RankNullity.lean | 12 +- .../Dimension/StrongRankCondition.lean | 8 ++ Mathlib/LinearAlgebra/Dual.lean | 11 ++ Mathlib/LinearAlgebra/RootSystem/Defs.lean | 6 + .../RootSystem/Finite/CanonicalBilinear.lean | 64 ++++++++- .../RootSystem/Finite/Nondegenerate.lean | 128 ++++++++++++++++++ docs/references.bib | 15 ++ 10 files changed, 257 insertions(+), 9 deletions(-) create mode 100644 Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean diff --git a/Mathlib.lean b/Mathlib.lean index 14b5b0c4d7f7a..5922aa04576e9 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3391,6 +3391,7 @@ import Mathlib.LinearAlgebra.Reflection import Mathlib.LinearAlgebra.RootSystem.Basic import Mathlib.LinearAlgebra.RootSystem.Defs import Mathlib.LinearAlgebra.RootSystem.Finite.CanonicalBilinear +import Mathlib.LinearAlgebra.RootSystem.Finite.Nondegenerate import Mathlib.LinearAlgebra.RootSystem.Hom import Mathlib.LinearAlgebra.RootSystem.OfBilinear import Mathlib.LinearAlgebra.RootSystem.RootPairingCat diff --git a/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean b/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean index 4bfa3ea69c30d..b51accd2d82ad 100644 --- a/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean @@ -95,7 +95,6 @@ lemma sum_sq_le_sq_sum_of_nonneg (hf : ∀ i ∈ s, 0 ≤ f i) : · exact hf i hi · exact single_le_sum hf hi - end OrderedSemiring section OrderedCommSemiring @@ -142,6 +141,19 @@ lemma sum_mul_sq_le_sq_mul_sq (s : Finset ι) (f g : ι → R) : sum_le_sum fun i _ ↦ two_mul_le_add_sq (f i * ∑ j ∈ s, g j ^ 2) (g i * ∑ j ∈ s, f j * g j) _ = _ := by simp_rw [sum_add_distrib, mul_pow, ← sum_mul]; ring +theorem sum_mul_self_eq_zero_iff (s : Finset ι) (f : ι → R) : + ∑ i ∈ s, f i * f i = 0 ↔ ∀ i ∈ s, f i = 0 := by + induction s using Finset.cons_induction with + | empty => simp + | cons i s his ih => + simp only [Finset.sum_cons, Finset.mem_cons, forall_eq_or_imp] + refine ⟨fun hc => ?_, fun h => by simpa [h.1] using ih.mpr h.2⟩ + have hi : f i * f i ≤ 0 := by + rw [← hc, le_add_iff_nonneg_right] + exact Finset.sum_nonneg fun i _ ↦ mul_self_nonneg (f i) + have h : f i * f i = 0 := (eq_of_le_of_le (mul_self_nonneg (f i)) hi).symm + exact ⟨zero_eq_mul_self.mp h.symm, ih.mp (by rw [← hc, h, zero_add])⟩ + end LinearOrderedCommSemiring lemma abs_prod [LinearOrderedCommRing R] (s : Finset ι) (f : ι → R) : diff --git a/Mathlib/LinearAlgebra/Dimension/Basic.lean b/Mathlib/LinearAlgebra/Dimension/Basic.lean index 76b2ee39e5e51..254ddee6fae30 100644 --- a/Mathlib/LinearAlgebra/Dimension/Basic.lean +++ b/Mathlib/LinearAlgebra/Dimension/Basic.lean @@ -353,6 +353,13 @@ theorem rank_subsingleton [Subsingleton R] : Module.rank R M = 1 := by subsingleton · exact hw.trans_eq (Cardinal.mk_singleton _).symm +lemma rank_le_of_smul_regular {S : Type*} [CommSemiring S] [Algebra S R] [Module S M] + [IsScalarTower S R M] (L L' : Submodule R M) {s : S} (hr : IsSMulRegular M s) + (h : ∀ x ∈ L, s • x ∈ L') : + Module.rank R L ≤ Module.rank R L' := + ((Algebra.lsmul S R M s).restrict h).rank_le_of_injective <| + fun _ _ h ↦ by simpa using hr (Subtype.ext_iff.mp h) + end end Module diff --git a/Mathlib/LinearAlgebra/Dimension/RankNullity.lean b/Mathlib/LinearAlgebra/Dimension/RankNullity.lean index 79f9779cf6271..cf3c16a79daa0 100644 --- a/Mathlib/LinearAlgebra/Dimension/RankNullity.lean +++ b/Mathlib/LinearAlgebra/Dimension/RankNullity.lean @@ -203,13 +203,23 @@ lemma Submodule.finrank_quotient_add_finrank [Module.Finite R M] (N : Submodule Submodule.finrank_eq_rank] exact HasRankNullity.rank_quotient_add_rank _ - /-- Rank-nullity theorem using `finrank` and subtraction. -/ lemma Submodule.finrank_quotient [Module.Finite R M] {S : Type*} [Ring S] [SMul R S] [Module S M] [IsScalarTower R S M] (N : Submodule S M) : finrank R (M ⧸ N) = finrank R M - finrank R N := by rw [← (N.restrictScalars R).finrank_quotient_add_finrank] exact Nat.eq_sub_of_add_eq rfl +lemma Submodule.disjoint_ker_of_finrank_eq [NoZeroSMulDivisors R M] {N : Type*} [AddCommGroup N] + [Module R N] {L : Submodule R M} [Module.Finite R L] (f : M →ₗ[R] N) + (h : finrank R (L.map f) = finrank R L) : + Disjoint L (LinearMap.ker f) := by + refine disjoint_iff.mpr <| LinearMap.injective_domRestrict_iff.mp <| LinearMap.ker_eq_bot.mp <| + Submodule.rank_eq_zero.mp ?_ + rw [← Submodule.finrank_eq_rank, Nat.cast_eq_zero] + rw [← LinearMap.range_domRestrict] at h + have := (LinearMap.ker (f.domRestrict L)).finrank_quotient_add_finrank + rwa [LinearEquiv.finrank_eq (f.domRestrict L).quotKerEquivRange, h, Nat.add_eq_left] at this + end Finrank section diff --git a/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean b/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean index cd51f9aec9d87..0e7836ca78bb4 100644 --- a/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean +++ b/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean @@ -465,4 +465,12 @@ theorem LinearMap.finrank_range_le [Module.Finite R M] (f : M →ₗ[R] M') : finrank R (LinearMap.range f) ≤ finrank R M := finrank_le_finrank_of_rank_le_rank (lift_rank_range_le f) (rank_lt_aleph0 _ _) +theorem LinearMap.finrank_le_of_smul_regular {S : Type*} [CommSemiring S] [Algebra S R] [Module S M] + [IsScalarTower S R M] (L L' : Submodule R M) [Module.Finite R L'] {s : S} + (hr : IsSMulRegular M s) (h : ∀ x ∈ L, s • x ∈ L') : + Module.finrank R L ≤ Module.finrank R L' := by + refine finrank_le_finrank_of_rank_le_rank (lift_le.mpr <| rank_le_of_smul_regular L L' hr h) ?_ + rw [← Module.finrank_eq_rank R L'] + exact nat_lt_aleph0 (finrank R ↥L') + end StrongRankCondition diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index c41abe513c0a7..1ef8aede9c87b 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -715,6 +715,17 @@ instance instFiniteDimensionalOfIsReflexive (K V : Type*) rw [lift_id'] exact lt_trans h₁ h₂ +instance [IsDomain R] : NoZeroSMulDivisors R M := by + refine (noZeroSMulDivisors_iff R M).mpr ?_ + intro r m hrm + rw [or_iff_not_imp_left] + intro hr + suffices Dual.eval R M m = Dual.eval R M 0 from (bijective_dual_eval R M).injective this + ext n + simp only [Dual.eval_apply, map_zero, LinearMap.zero_apply] + suffices r • n m = 0 from eq_zero_of_ne_zero_of_mul_left_eq_zero hr this + rw [← LinearMap.map_smul_of_tower, hrm, LinearMap.map_zero] + end IsReflexive end Module diff --git a/Mathlib/LinearAlgebra/RootSystem/Defs.lean b/Mathlib/LinearAlgebra/RootSystem/Defs.lean index 420de4b340130..d9ab0eb932455 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Defs.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Defs.lean @@ -374,6 +374,12 @@ lemma isReduced_iff : P.IsReduced ↔ ∀ i j : ι, i ≠ j → · exact Or.inl (congrArg P.root h') · exact Or.inr (h i j h' hLin) +/-- The linear span of roots. -/ +abbrev rootSpan := span R (range P.root) + +/-- The linear span of coroots. -/ +abbrev corootSpan := span R (range P.coroot) + /-- The `Weyl group` of a root pairing is the group of automorphisms of the weight space generated by reflections in roots. -/ def weylGroup : Subgroup (M ≃ₗ[R] M) := diff --git a/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean b/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean index d495f1b08bbcc..1f7676b71d4ad 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean @@ -3,8 +3,9 @@ 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.Order.BigOperators.Ring.Finset import Mathlib.Algebra.Ring.SumsOfSquares +import Mathlib.LinearAlgebra.RootSystem.Defs /-! # The canonical bilinear form on a finite root pairing @@ -22,27 +23,28 @@ Weyl group. * `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. + * `rootForm_self_non_neg`: `RootForm` is positive semidefinite. + +## References: + * [N. Bourbaki, *Lie groups and {L}ie algebras. {C}hapters 4--6*][bourbaki1968] + * [M. Demazure, *SGA III, Expos\'{e} XXI, Don\'{e}es Radicielles*][demazure1970] ## 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 Set Function open Module hiding reflection +open Submodule (span) noncomputable section @@ -132,6 +134,54 @@ 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] +theorem range_polarization_domRestrict_le_span_coroot : + LinearMap.range (P.Polarization.domRestrict P.rootSpan) ≤ P.corootSpan := by + intro y hy + obtain ⟨x, hx⟩ := hy + rw [← hx, LinearMap.domRestrict_apply, Polarization_apply] + refine (mem_span_range_iff_exists_fun R).mpr ?_ + use fun i => (P.toPerfectPairing x) (P.coroot i) + simp + end CommRing +section LinearOrderedCommRing + +variable [Fintype ι] [LinearOrderedCommRing R] [AddCommGroup M] [Module R M] [AddCommGroup N] + [Module R N] (P : RootPairing ι R M N) + +theorem rootForm_self_non_neg (x : M) : 0 ≤ P.RootForm x x := + IsSumSq.nonneg (P.rootForm_self_sum_of_squares x) + +theorem rootForm_self_zero_iff (x : M) : + P.RootForm x x = 0 ↔ ∀ i, P.coroot' i x = 0 := by + simp only [rootForm_apply_apply, PerfectPairing.toLin_apply, LinearMap.coe_comp, comp_apply, + Polarization_apply, map_sum, map_smul, smul_eq_mul] + convert Finset.sum_mul_self_eq_zero_iff Finset.univ fun i => P.coroot' i x + simp + +lemma rootForm_root_self_pos (j : ι) : + 0 < P.RootForm (P.root j) (P.root j) := by + simp only [LinearMap.coe_mk, AddHom.coe_mk, LinearMap.coe_comp, comp_apply, + rootForm_apply_apply, toLin_toPerfectPairing] + refine Finset.sum_pos' (fun i _ => (sq (P.pairing j i)) ▸ sq_nonneg (P.pairing j i)) ?_ + use j + simp + +lemma prod_rootForm_root_self_pos : + 0 < ∏ i, P.RootForm (P.root i) (P.root i) := + Finset.prod_pos fun i _ => rootForm_root_self_pos P i + +lemma prod_rootForm_smul_coroot_mem_range_domRestrict (i : ι) : + (∏ a : ι, P.RootForm (P.root a) (P.root a)) • P.coroot i ∈ + LinearMap.range (P.Polarization.domRestrict (P.rootSpan)) := by + obtain ⟨c, hc⟩ := Finset.dvd_prod_of_mem (fun a ↦ P.RootForm (P.root a) (P.root a)) + (Finset.mem_univ i) + rw [hc, mul_comm, mul_smul, rootForm_self_smul_coroot] + refine LinearMap.mem_range.mpr ?_ + use ⟨(c • 2 • P.root i), by aesop⟩ + simp + +end LinearOrderedCommRing + end RootPairing diff --git a/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean b/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean new file mode 100644 index 0000000000000..734860e88627e --- /dev/null +++ b/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean @@ -0,0 +1,128 @@ +/- +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.BilinearForm.Basic +import Mathlib.LinearAlgebra.Dimension.Localization +import Mathlib.LinearAlgebra.QuadraticForm.Basic +import Mathlib.LinearAlgebra.RootSystem.Finite.CanonicalBilinear +import Mathlib.LinearAlgebra.RootSystem.RootPositive + +/-! +# Nondegeneracy of the polarization on a finite root pairing + +We show that if the base ring of a finite root pairing is linearly ordered, then the canonical +bilinear form is root-positive 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 root pairing are restricted to the +interval `[-4, 4]`. Furthermore, a linearly independent pair of roots cannot have Coxeter weight 4. +For the case of crystallographic root pairings, we are thus reduced to a finite set of possible +options for each pair. +Another application is to the faithfulness of the Weyl group action on roots, and finiteness of the +Weyl group. + +## Main results: + * `RootPairing.rootForm_rootPositive`: `RootForm` is root-positive. + * `RootPairing.polarization_domRestrict_injective`: The polarization restricted to the span of + roots is injective. + * `RootPairing.rootForm_pos_of_nonzero`: `RootForm` is strictly positive on non-zero linear + combinations of roots. This gives us a convenient way to eliminate certain Dynkin diagrams from + the classification, since it suffices to produce a nonzero linear combination of simple roots with + non-positive norm. + +## References: + * [N. Bourbaki, *Lie groups and {L}ie algebras. {C}hapters 4--6*][bourbaki1968] + * [M. Demazure, *SGA III, Expos\'{e} XXI, Don\'{e}es Radicielles*][demazure1970] + +## Todo + * Weyl-invariance of `RootForm` and `CorootForm` + * Faithfulness of Weyl group perm action, and finiteness of Weyl group, over ordered rings. + * Relation to Coxeter weight. In particular, positivity constraints for finite root pairings mean + we restrict to weights between 0 and 4. +-/ + +noncomputable section + +open Set Function +open Module hiding reflection +open Submodule (span) + +namespace RootPairing + +variable {ι R M N : Type*} + +variable [Fintype ι] [LinearOrderedCommRing R] [AddCommGroup M] [Module R M] [AddCommGroup N] +[Module R N] (P : RootPairing ι R M N) + +lemma rootForm_rootPositive : IsRootPositive P P.RootForm where + zero_lt_apply_root i := P.rootForm_root_self_pos i + symm := P.rootForm_symmetric + apply_reflection_eq := P.rootForm_reflection_reflection_apply + +instance : Module.Finite R P.rootSpan := Finite.span_of_finite R <| finite_range P.root + +instance : Module.Finite R P.corootSpan := Finite.span_of_finite R <| finite_range P.coroot + +@[simp] +lemma finrank_rootSpan_map_polarization_eq_finrank_corootSpan : + finrank R (P.rootSpan.map P.Polarization) = finrank R P.corootSpan := by + rw [← LinearMap.range_domRestrict] + apply (Submodule.finrank_mono P.range_polarization_domRestrict_le_span_coroot).antisymm + have : IsReflexive R N := PerfectPairing.reflexive_right P.toPerfectPairing + refine LinearMap.finrank_le_of_smul_regular P.corootSpan + (LinearMap.range (P.Polarization.domRestrict P.rootSpan)) + (smul_right_injective N (Ne.symm (ne_of_lt P.prod_rootForm_root_self_pos))) + fun _ hx => ?_ + obtain ⟨c, hc⟩ := (mem_span_range_iff_exists_fun R).mp hx + rw [← hc, Finset.smul_sum] + simp_rw [smul_smul, mul_comm, ← smul_smul] + exact Submodule.sum_smul_mem (LinearMap.range (P.Polarization.domRestrict P.rootSpan)) c + (fun c _ ↦ prod_rootForm_smul_coroot_mem_range_domRestrict P c) + +/-- An auxiliary lemma en route to `RootPairing.finrank_corootSpan_eq`. -/ +private lemma finrank_corootSpan_le : + finrank R P.corootSpan ≤ finrank R P.rootSpan := by + rw [← finrank_rootSpan_map_polarization_eq_finrank_corootSpan] + exact Submodule.finrank_map_le P.Polarization P.rootSpan + +lemma finrank_corootSpan_eq : + finrank R P.corootSpan = finrank R P.rootSpan := + le_antisymm P.finrank_corootSpan_le P.flip.finrank_corootSpan_le + +lemma disjoint_rootSpan_ker_polarization : + Disjoint P.rootSpan (LinearMap.ker P.Polarization) := by + have : IsReflexive R M := PerfectPairing.reflexive_left P.toPerfectPairing + refine Submodule.disjoint_ker_of_finrank_eq (L := P.rootSpan) P.Polarization ?_ + rw [finrank_rootSpan_map_polarization_eq_finrank_corootSpan, finrank_corootSpan_eq] + +lemma mem_ker_polarization_of_rootForm_self_eq_zero {x : M} (hx : P.RootForm x x = 0) : + x ∈ LinearMap.ker P.Polarization := by + rw [LinearMap.mem_ker, Polarization_apply] + rw [rootForm_self_zero_iff] at hx + exact Fintype.sum_eq_zero _ fun i ↦ by simp [hx i] + +lemma eq_zero_of_mem_rootSpan_of_rootForm_self_eq_zero {x : M} + (hx : x ∈ P.rootSpan) (hx' : P.RootForm x x = 0) : + x = 0 := by + rw [← Submodule.mem_bot (R := R), ← P.disjoint_rootSpan_ker_polarization.eq_bot] + exact ⟨hx, P.mem_ker_polarization_of_rootForm_self_eq_zero hx'⟩ + +lemma _root_.RootSystem.rootForm_anisotropic (P : RootSystem ι R M N) : + P.RootForm.toQuadraticMap.Anisotropic := + fun x ↦ P.eq_zero_of_mem_rootSpan_of_rootForm_self_eq_zero <| by + simpa only [rootSpan, P.span_eq_top] using Submodule.mem_top + +lemma rootForm_pos_of_nonzero {x : M} (hx : x ∈ P.rootSpan) (h : x ≠ 0) : + 0 < P.RootForm x x := by + apply (P.rootForm_self_non_neg x).lt_of_ne + contrapose! h + exact eq_zero_of_mem_rootSpan_of_rootForm_self_eq_zero P hx h.symm + +lemma rootForm_restrict_nondegenerate : + (P.RootForm.restrict P.rootSpan).Nondegenerate := + LinearMap.IsRefl.nondegenerate_of_separatingLeft (LinearMap.IsSymm.isRefl fun x y => by + simp [rootForm_apply_apply, mul_comm]) fun x h => SetLike.coe_eq_coe.mp + (P.eq_zero_of_mem_rootSpan_of_rootForm_self_eq_zero (Submodule.coe_mem x) (h x)) + +end RootPairing diff --git a/docs/references.bib b/docs/references.bib index 8d65ce04a4a4e..6603b6a5466fe 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -1090,6 +1090,21 @@ @InProceedings{ deligne_formulaire mrreviewer = {Jacques Velu} } +@InProceedings{ demazure1970, + author = {Michel Demazure}, + editor = {M. Demazure, A. Grothendieck}, + title = {Expos\'{e} XXI, Don\'{e}es Radicielles}, + booktitle = {S\'{e}minaire de G\'{e}ométrie Alg\'{e}brique du Bois + Marie - 1962-64 - Sch\'{e}mas en groupes - (SGA 3) - vol. + 3, Structure des Sch\'{e}mas en Groupes Reductifs}, + series = {Lecture Notes in Mathematics}, + volume = {153}, + pages = {85--155}, + publisher = {Springer-Verlag}, + year = {1970}, + url = {https://wstein.org/sga/SGA3/Expo21-alpha.pdf} +} + @InProceedings{ demoura2015lean, author = {de Moura, Leonardo and Kong, Soonho and Avigad, Jeremy and van Doorn, Floris and von Raumer, Jakob}, From 961ee9dbff2b14a178fa757f097b5bb27c274e99 Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 20 Nov 2024 20:59:58 +0000 Subject: [PATCH 163/829] fix: more explicit PR number in `bors x` actions (#19304) --- .github/workflows/maintainer_bors.yml | 7 +++++-- 1 file changed, 5 insertions(+), 2 deletions(-) diff --git a/.github/workflows/maintainer_bors.yml b/.github/workflows/maintainer_bors.yml index 6832b6e007eed..073477b4fa324 100644 --- a/.github/workflows/maintainer_bors.yml +++ b/.github/workflows/maintainer_bors.yml @@ -43,6 +43,9 @@ jobs: printf $'"bors delegate" or "bors merge" found? \'%s\'\n' "${m_or_d}" printf $'AUTHOR: \'%s\'\n' "${AUTHOR}" + printf $'PR_NUMBER: \'%s\'\n' "${PR_NUMBER}" + printf $'OTHER_NUMBER: \'%s\'\n' "${{ github.event.issue.number }}${{ github.event.pull_request.number }}" + printf $'%s' "${{ github.event.issue.number }}${{ github.event.pull_request.number }}" | hexdump -cC printf $'mOrD=%s\n' "${m_or_d}" >> "${GITHUB_OUTPUT}" if [ "${AUTHOR}" == 'leanprover-community-mathlib4-bot' ] || @@ -73,9 +76,9 @@ jobs: uses: octokit/request-action@v2.x with: # check is this ok? was /repos/:repository/issues/:issue_number/labels - route: POST /repos/:repository/issues/${PR_NUMBER}/labels + route: POST /repos/:repository/issues/:issue_number/labels repository: ${{ github.repository }} - issue_number: ${PR_NUMBER} + issue_number: ${{ github.event.issue.number }}${{ github.event.pull_request.number }} labels: '["${{ steps.merge_or_delegate.outputs.mOrD }}"]' env: GITHUB_TOKEN: ${{ secrets.TRIAGE_TOKEN }} # comment uses ${{ secrets.GITHUB_TOKEN }} From c44399da9f3d7eab24f365b684f7f1cb18adb8bc Mon Sep 17 00:00:00 2001 From: Pietro Monticone <38562595+pitmonticone@users.noreply.github.com> Date: Wed, 20 Nov 2024 22:00:50 +0000 Subject: [PATCH 164/829] feat(Order/Max): add `IsTop.isMax_iff` and `IsBot.isMin_iff` (#19305) Upstreamed from the [Carleson](https://github.com/fpvandoorn/carleson) project. Co-authored-by: Floris van Doorn --- Mathlib/Order/Max.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/Mathlib/Order/Max.lean b/Mathlib/Order/Max.lean index 36518a3cb4aa5..7cf8fc7da8e5c 100644 --- a/Mathlib/Order/Max.lean +++ b/Mathlib/Order/Max.lean @@ -200,6 +200,14 @@ protected theorem IsBot.isMin (h : IsBot a) : IsMin a := fun b _ => h b protected theorem IsTop.isMax (h : IsTop a) : IsMax a := fun b _ => h b +theorem IsTop.isMax_iff {α} [PartialOrder α] {i j : α} (h : IsTop i) : IsMax j ↔ j = i := by + simp_rw [le_antisymm_iff, h j, true_and] + exact ⟨(· (h j)), Function.swap (fun _ ↦ h · |>.trans ·)⟩ + +theorem IsBot.isMin_iff {α} [PartialOrder α] {i j : α} (h : IsBot i) : IsMin j ↔ j = i := by + simp_rw [le_antisymm_iff, h j, and_true] + exact ⟨fun a ↦ a (h j), fun a h' ↦ fun _ ↦ Preorder.le_trans j i h' a (h h')⟩ + @[simp] theorem isBot_toDual_iff : IsBot (toDual a) ↔ IsTop a := Iff.rfl From 931c0de5d6b5105b720e87c95de99c5a710b69cd Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Wed, 20 Nov 2024 22:27:27 +0000 Subject: [PATCH 165/829] chore: reducibility in gcongr apply step (#19262) [Zulip](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/gcongr.20unfold.20DFunLike.2Ecoe) --- Mathlib/Tactic/GCongr/Core.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Tactic/GCongr/Core.lean b/Mathlib/Tactic/GCongr/Core.lean index a47257466df27..90aeed7ba3815 100644 --- a/Mathlib/Tactic/GCongr/Core.lean +++ b/Mathlib/Tactic/GCongr/Core.lean @@ -367,7 +367,7 @@ partial def _root_.Lean.MVarId.gcongr for lem in (gcongrExt.getState (← getEnv)).getD (relName, lhsHead, varyingArgs) #[] do let gs ← try -- Try `apply`-ing such a lemma to the goal. - Except.ok <$> g.apply (← mkConstWithFreshMVarLevels lem.declName) + Except.ok <$> withReducibleAndInstances (g.apply (← mkConstWithFreshMVarLevels lem.declName)) catch e => pure (Except.error e) match gs with | .error e => From 232f84f52613e3ad9e60ceed1f57841e6de39fb7 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Wed, 20 Nov 2024 22:27:28 +0000 Subject: [PATCH 166/829] chore: nolint `div_le_div_iff'` (#19309) This lemma was accidentally removed from the nolints file in #18917, then accidentally had its warning silenced in-place in #19279; this PR undoes both changes! --- Mathlib/Algebra/Order/Group/Unbundled/Basic.lean | 1 - scripts/no_lints_prime_decls.txt | 1 + 2 files changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean index 5b90caf6e8964..03c9d7579e5f1 100644 --- a/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean @@ -537,7 +537,6 @@ section LE variable [LE α] [MulLeftMono α] {a b c d : α} -set_option linter.docPrime false in @[to_additive sub_le_sub_iff] theorem div_le_div_iff' : a / b ≤ c / d ↔ a * d ≤ c * b := by simpa only [div_eq_mul_inv] using mul_inv_le_mul_inv_iff' diff --git a/scripts/no_lints_prime_decls.txt b/scripts/no_lints_prime_decls.txt index 7a7e8fb75bfe5..e04b49b8dc0f0 100644 --- a/scripts/no_lints_prime_decls.txt +++ b/scripts/no_lints_prime_decls.txt @@ -1088,6 +1088,7 @@ div_eq_iff_eq_mul' div_eq_of_eq_mul' div_eq_of_eq_mul'' div_le_div'' +div_le_div_iff' div_le_div_left' div_le_div_right' div_left_inj' From 50c8065ea373f936358889d6eae523a8fde8ae34 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 20 Nov 2024 23:50:35 +0000 Subject: [PATCH 167/829] chore: import `Tactic.Positivity.Finset` in `Tactic.Positivity` (#18146) I keep on getting bitten by the missing import. --- Mathlib/Algebra/AddConstMap/Basic.lean | 2 ++ Mathlib/Algebra/AddConstMap/Equiv.lean | 3 +++ Mathlib/Algebra/ContinuedFractions/Computation/Basic.lean | 1 + .../Algebra/ContinuedFractions/Computation/Translations.lean | 1 + Mathlib/Algebra/Order/Archimedean/Basic.lean | 2 ++ Mathlib/Algebra/Order/Archimedean/Hom.lean | 2 ++ Mathlib/Algebra/Order/Archimedean/Submonoid.lean | 2 ++ Mathlib/Algebra/Order/Chebyshev.lean | 3 +-- Mathlib/Algebra/Order/Floor.lean | 4 +++- Mathlib/Algebra/Order/Nonneg/Floor.lean | 2 ++ Mathlib/Analysis/Analytic/Inverse.lean | 2 +- Mathlib/Analysis/Complex/AbelLimit.lean | 2 +- Mathlib/Analysis/Normed/Group/Seminorm.lean | 1 + Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean | 1 - Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean | 1 - Mathlib/Data/ENNReal/Basic.lean | 1 + Mathlib/Data/Finset/Density.lean | 2 +- Mathlib/Data/Int/Log.lean | 1 + Mathlib/Data/Int/WithZero.lean | 2 ++ Mathlib/Data/NNRat/Floor.lean | 2 ++ Mathlib/Data/NNReal/Star.lean | 2 ++ Mathlib/Data/Rat/Floor.lean | 1 + Mathlib/Data/Real/Archimedean.lean | 2 ++ Mathlib/Data/Real/ENatENNReal.lean | 1 + Mathlib/Data/Real/Pointwise.lean | 1 + Mathlib/GroupTheory/Archimedean.lean | 2 ++ Mathlib/LinearAlgebra/Ray.lean | 2 +- Mathlib/NumberTheory/Harmonic/Defs.lean | 2 +- Mathlib/Tactic/Positivity.lean | 1 + MathlibTest/positivity.lean | 2 +- 30 files changed, 42 insertions(+), 11 deletions(-) diff --git a/Mathlib/Algebra/AddConstMap/Basic.lean b/Mathlib/Algebra/AddConstMap/Basic.lean index 7dc43d9c08e9a..b8a261ee2098e 100644 --- a/Mathlib/Algebra/AddConstMap/Basic.lean +++ b/Mathlib/Algebra/AddConstMap/Basic.lean @@ -28,6 +28,8 @@ We use parameters `a` and `b` instead of `1` to accommodate for two use cases: including orientation-reversing maps. -/ +assert_not_exists Finset + open Function Set /-- A bundled map `f : G → H` such that `f (x + a) = f x + b` for all `x`. diff --git a/Mathlib/Algebra/AddConstMap/Equiv.lean b/Mathlib/Algebra/AddConstMap/Equiv.lean index ee4c37b716812..28779a4f482cc 100644 --- a/Mathlib/Algebra/AddConstMap/Equiv.lean +++ b/Mathlib/Algebra/AddConstMap/Equiv.lean @@ -5,6 +5,7 @@ Authors: Yury Kudryashov -/ import Mathlib.Algebra.AddConstMap.Basic import Mathlib.GroupTheory.Perm.Basic + /-! # Equivalences conjugating `(· + a)` to `(· + b)` @@ -14,6 +15,8 @@ to be the type of equivalences such that `∀ x, f (x + a) = f x + b`. We also define the corresponding typeclass and prove some basic properties. -/ +assert_not_exists Finset + open Function open scoped AddConstMap diff --git a/Mathlib/Algebra/ContinuedFractions/Computation/Basic.lean b/Mathlib/Algebra/ContinuedFractions/Computation/Basic.lean index 44e1c3d71de18..f269dd69e8670 100644 --- a/Mathlib/Algebra/ContinuedFractions/Computation/Basic.lean +++ b/Mathlib/Algebra/ContinuedFractions/Computation/Basic.lean @@ -58,6 +58,7 @@ with a head term (`seq1`) is then transformed to a generalized continued fractio numerics, number theory, approximations, fractions -/ +assert_not_exists Finset namespace GenContFract diff --git a/Mathlib/Algebra/ContinuedFractions/Computation/Translations.lean b/Mathlib/Algebra/ContinuedFractions/Computation/Translations.lean index d732ff25b6a7a..0b971d4795163 100644 --- a/Mathlib/Algebra/ContinuedFractions/Computation/Translations.lean +++ b/Mathlib/Algebra/ContinuedFractions/Computation/Translations.lean @@ -37,6 +37,7 @@ of three sections: parts. -/ +assert_not_exists Finset namespace GenContFract diff --git a/Mathlib/Algebra/Order/Archimedean/Basic.lean b/Mathlib/Algebra/Order/Archimedean/Basic.lean index 7756989822b78..c2c50a58878f6 100644 --- a/Mathlib/Algebra/Order/Archimedean/Basic.lean +++ b/Mathlib/Algebra/Order/Archimedean/Basic.lean @@ -30,6 +30,8 @@ number `n` such that `x ≤ n • y`. * `ℕ`, `ℤ`, and `ℚ` are archimedean. -/ +assert_not_exists Finset + open Int Set variable {α : Type*} diff --git a/Mathlib/Algebra/Order/Archimedean/Hom.lean b/Mathlib/Algebra/Order/Archimedean/Hom.lean index c4e7b11c89739..03816d8b0f244 100644 --- a/Mathlib/Algebra/Order/Archimedean/Hom.lean +++ b/Mathlib/Algebra/Order/Archimedean/Hom.lean @@ -14,6 +14,8 @@ ordered field. Reciprocally, such an ordered ring homomorphism exists when the c conditionally complete. -/ +assert_not_exists Finset + variable {α β : Type*} /-- There is at most one ordered ring homomorphism from a linear ordered field to an archimedean diff --git a/Mathlib/Algebra/Order/Archimedean/Submonoid.lean b/Mathlib/Algebra/Order/Archimedean/Submonoid.lean index fccc02aa01ab0..b30d6bec3e7ff 100644 --- a/Mathlib/Algebra/Order/Archimedean/Submonoid.lean +++ b/Mathlib/Algebra/Order/Archimedean/Submonoid.lean @@ -22,6 +22,8 @@ submonoid of the ambient group. submonoid. -/ +assert_not_exists Finset + @[to_additive] instance SubmonoidClass.instMulArchimedean {M S : Type*} [SetLike S M] [OrderedCommMonoid M] [SubmonoidClass S M] [MulArchimedean M] (H : S) : MulArchimedean H := by diff --git a/Mathlib/Algebra/Order/Chebyshev.lean b/Mathlib/Algebra/Order/Chebyshev.lean index b0be97b492f03..d359503ea08c9 100644 --- a/Mathlib/Algebra/Order/Chebyshev.lean +++ b/Mathlib/Algebra/Order/Chebyshev.lean @@ -7,8 +7,7 @@ import Mathlib.Algebra.Order.Monovary import Mathlib.Algebra.Order.Rearrangement import Mathlib.GroupTheory.Perm.Cycle.Basic import Mathlib.Tactic.GCongr -import Mathlib.Tactic.Positivity.Basic -import Mathlib.Tactic.Positivity.Finset +import Mathlib.Tactic.Positivity /-! # Chebyshev's sum inequality diff --git a/Mathlib/Algebra/Order/Floor.lean b/Mathlib/Algebra/Order/Floor.lean index 5e5aaf78be137..d3d6e7b16dc1d 100644 --- a/Mathlib/Algebra/Order/Floor.lean +++ b/Mathlib/Algebra/Order/Floor.lean @@ -13,7 +13,7 @@ import Mathlib.Order.GaloisConnection import Mathlib.Tactic.Abel import Mathlib.Tactic.FieldSimp import Mathlib.Tactic.Linarith -import Mathlib.Tactic.Positivity +import Mathlib.Tactic.Positivity.Basic /-! # Floor and ceil @@ -54,6 +54,8 @@ many lemmas. rounding, floor, ceil -/ +assert_not_exists Finset + open Set variable {F α β : Type*} diff --git a/Mathlib/Algebra/Order/Nonneg/Floor.lean b/Mathlib/Algebra/Order/Nonneg/Floor.lean index a3472c9dde95b..1fdd9eefd1892 100644 --- a/Mathlib/Algebra/Order/Nonneg/Floor.lean +++ b/Mathlib/Algebra/Order/Nonneg/Floor.lean @@ -19,6 +19,8 @@ This is used to derive algebraic structures on `ℝ≥0` and `ℚ≥0` automatic * `{x : α // 0 ≤ x}` is a `FloorSemiring` if `α` is. -/ +assert_not_exists Finset + namespace Nonneg variable {α : Type*} diff --git a/Mathlib/Analysis/Analytic/Inverse.lean b/Mathlib/Analysis/Analytic/Inverse.lean index c9a57e5bb6985..6efcaa447e727 100644 --- a/Mathlib/Analysis/Analytic/Inverse.lean +++ b/Mathlib/Analysis/Analytic/Inverse.lean @@ -5,7 +5,7 @@ Authors: Sébastien Gouëzel -/ import Mathlib.Analysis.Analytic.Composition import Mathlib.Analysis.Analytic.Linear -import Mathlib.Tactic.Positivity.Finset +import Mathlib.Tactic.Positivity /-! diff --git a/Mathlib/Analysis/Complex/AbelLimit.lean b/Mathlib/Analysis/Complex/AbelLimit.lean index 25d7f649514c6..7e5aea15a8967 100644 --- a/Mathlib/Analysis/Complex/AbelLimit.lean +++ b/Mathlib/Analysis/Complex/AbelLimit.lean @@ -6,7 +6,7 @@ Authors: Jeremy Tan import Mathlib.Analysis.Complex.Basic import Mathlib.Analysis.SpecificLimits.Normed import Mathlib.Tactic.Peel -import Mathlib.Tactic.Positivity.Finset +import Mathlib.Tactic.Positivity /-! # Abel's limit theorem diff --git a/Mathlib/Analysis/Normed/Group/Seminorm.lean b/Mathlib/Analysis/Normed/Group/Seminorm.lean index 38c8d7fee64ca..e895bf01b8d3c 100644 --- a/Mathlib/Analysis/Normed/Group/Seminorm.lean +++ b/Mathlib/Analysis/Normed/Group/Seminorm.lean @@ -44,6 +44,7 @@ having a superfluous `add_le'` field in the resulting structure. The same applie norm, seminorm -/ +assert_not_exists Finset open Set diff --git a/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean b/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean index a65c2df01ec20..0a386abf7a0f9 100644 --- a/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean +++ b/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean @@ -10,7 +10,6 @@ import Mathlib.Combinatorics.Enumerative.DoubleCounting import Mathlib.Tactic.FieldSimp import Mathlib.Tactic.GCongr import Mathlib.Tactic.Positivity -import Mathlib.Tactic.Positivity.Finset import Mathlib.Tactic.Ring /-! diff --git a/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean b/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean index a5d36c33b508b..6d675a7865092 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean @@ -10,7 +10,6 @@ import Mathlib.Combinatorics.SimpleGraph.Clique import Mathlib.Data.Finset.Sym import Mathlib.Tactic.GCongr import Mathlib.Tactic.Positivity -import Mathlib.Tactic.Positivity.Finset /-! # Triangles in graphs diff --git a/Mathlib/Data/ENNReal/Basic.lean b/Mathlib/Data/ENNReal/Basic.lean index 96836f66517c4..120674ea5f000 100644 --- a/Mathlib/Data/ENNReal/Basic.lean +++ b/Mathlib/Data/ENNReal/Basic.lean @@ -85,6 +85,7 @@ context, or if we have `(f : α → ℝ≥0∞) (hf : ∀ x, f x ≠ ∞)`. -/ +assert_not_exists Finset open Function Set NNReal diff --git a/Mathlib/Data/Finset/Density.lean b/Mathlib/Data/Finset/Density.lean index b8d5585f00ed8..29e19adc50a21 100644 --- a/Mathlib/Data/Finset/Density.lean +++ b/Mathlib/Data/Finset/Density.lean @@ -7,7 +7,7 @@ import Mathlib.Algebra.Order.Field.Rat import Mathlib.Data.Fintype.Card import Mathlib.Data.NNRat.Order import Mathlib.Data.Rat.Cast.CharZero -import Mathlib.Tactic.Positivity +import Mathlib.Tactic.Positivity.Basic /-! # Density of a finite set diff --git a/Mathlib/Data/Int/Log.lean b/Mathlib/Data/Int/Log.lean index e6f5f9f457159..4d5f3c0e38722 100644 --- a/Mathlib/Data/Int/Log.lean +++ b/Mathlib/Data/Int/Log.lean @@ -46,6 +46,7 @@ def digits (b : ℕ) (q : ℚ) (n : ℕ) : ℕ := * `Int.neg_log_inv_eq_clog`, `Int.neg_clog_inv_eq_log`: the link between the two definitions. -/ +assert_not_exists Finset variable {R : Type*} [LinearOrderedSemifield R] [FloorSemiring R] diff --git a/Mathlib/Data/Int/WithZero.lean b/Mathlib/Data/Int/WithZero.lean index 4a74559dbbf98..5cf5aebd6da66 100644 --- a/Mathlib/Data/Int/WithZero.lean +++ b/Mathlib/Data/Int/WithZero.lean @@ -26,6 +26,8 @@ the morphism `WithZeroMultInt.toNNReal`. WithZero, multiplicative, nnreal -/ +assert_not_exists Finset + noncomputable section open scoped NNReal diff --git a/Mathlib/Data/NNRat/Floor.lean b/Mathlib/Data/NNRat/Floor.lean index 9fffafbea6995..02d70c2c3d123 100644 --- a/Mathlib/Data/NNRat/Floor.lean +++ b/Mathlib/Data/NNRat/Floor.lean @@ -19,6 +19,8 @@ Note that we cannot talk about `Int.fract`, which currently only works for rings nnrat, rationals, ℚ≥0, floor -/ +assert_not_exists Finset + namespace NNRat instance : FloorSemiring ℚ≥0 where diff --git a/Mathlib/Data/NNReal/Star.lean b/Mathlib/Data/NNReal/Star.lean index c1f11b1539a4e..47c6a75c17849 100644 --- a/Mathlib/Data/NNReal/Star.lean +++ b/Mathlib/Data/NNReal/Star.lean @@ -10,6 +10,8 @@ import Mathlib.Data.Real.Star # The non-negative real numbers are a `*`-ring, with the trivial `*`-structure -/ +assert_not_exists Finset + open scoped NNReal instance : StarRing ℝ≥0 := starRingOfComm diff --git a/Mathlib/Data/Rat/Floor.lean b/Mathlib/Data/Rat/Floor.lean index 66006dd0fc062..7b67abd634887 100644 --- a/Mathlib/Data/Rat/Floor.lean +++ b/Mathlib/Data/Rat/Floor.lean @@ -21,6 +21,7 @@ division and modulo arithmetic are derived as well as some simple inequalities. rat, rationals, ℚ, floor -/ +assert_not_exists Finset open Int diff --git a/Mathlib/Data/Real/Archimedean.lean b/Mathlib/Data/Real/Archimedean.lean index 55226e28d84b8..a42f86a993916 100644 --- a/Mathlib/Data/Real/Archimedean.lean +++ b/Mathlib/Data/Real/Archimedean.lean @@ -14,6 +14,8 @@ import Mathlib.Order.Interval.Set.Disjoint -/ +assert_not_exists Finset + open scoped Classical open Pointwise CauSeq diff --git a/Mathlib/Data/Real/ENatENNReal.lean b/Mathlib/Data/Real/ENatENNReal.lean index 9385bd57ec211..f022dc59fe0b4 100644 --- a/Mathlib/Data/Real/ENatENNReal.lean +++ b/Mathlib/Data/Real/ENatENNReal.lean @@ -12,6 +12,7 @@ import Mathlib.Data.ENNReal.Basic In this file we define a coercion from `ℕ∞` to `ℝ≥0∞` and prove some basic lemmas about this map. -/ +assert_not_exists Finset open scoped Classical open NNReal ENNReal diff --git a/Mathlib/Data/Real/Pointwise.lean b/Mathlib/Data/Real/Pointwise.lean index 8672dab314d36..e071edd0c5ddc 100644 --- a/Mathlib/Data/Real/Pointwise.lean +++ b/Mathlib/Data/Real/Pointwise.lean @@ -22,6 +22,7 @@ This is true more generally for conditionally complete linear order whose defaul don't have those yet. -/ +assert_not_exists Finset open Set diff --git a/Mathlib/GroupTheory/Archimedean.lean b/Mathlib/GroupTheory/Archimedean.lean index c31b5072d4f20..e67010ff8ade4 100644 --- a/Mathlib/GroupTheory/Archimedean.lean +++ b/Mathlib/GroupTheory/Archimedean.lean @@ -31,6 +31,8 @@ The result is also used in `Topology.Instances.Real` as an ingredient in the cla subgroups of `ℝ`. -/ +assert_not_exists Finset + open Set variable {G : Type*} [LinearOrderedCommGroup G] [MulArchimedean G] diff --git a/Mathlib/LinearAlgebra/Ray.lean b/Mathlib/LinearAlgebra/Ray.lean index 7890952886f09..d569b5c080c98 100644 --- a/Mathlib/LinearAlgebra/Ray.lean +++ b/Mathlib/LinearAlgebra/Ray.lean @@ -6,7 +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 +import Mathlib.Tactic.Positivity.Basic /-! # Rays in modules diff --git a/Mathlib/NumberTheory/Harmonic/Defs.lean b/Mathlib/NumberTheory/Harmonic/Defs.lean index 51c8f7036e358..29d3f5fa6bbe5 100644 --- a/Mathlib/NumberTheory/Harmonic/Defs.lean +++ b/Mathlib/NumberTheory/Harmonic/Defs.lean @@ -5,7 +5,7 @@ Authors: Koundinya Vajjha, Thomas Browning -/ import Mathlib.Algebra.BigOperators.Intervals import Mathlib.Algebra.Order.Field.Basic -import Mathlib.Tactic.Positivity.Finset +import Mathlib.Tactic.Positivity /-! diff --git a/Mathlib/Tactic/Positivity.lean b/Mathlib/Tactic/Positivity.lean index 8d76b597bcf25..7b2a24e70ec16 100644 --- a/Mathlib/Tactic/Positivity.lean +++ b/Mathlib/Tactic/Positivity.lean @@ -1,3 +1,4 @@ import Mathlib.Tactic.Positivity.Basic +import Mathlib.Tactic.Positivity.Finset import Mathlib.Tactic.NormNum.Basic import Mathlib.Data.Int.Order.Basic diff --git a/MathlibTest/positivity.lean b/MathlibTest/positivity.lean index da296d8d2860f..b3f2b9b7af6ba 100644 --- a/MathlibTest/positivity.lean +++ b/MathlibTest/positivity.lean @@ -1,10 +1,10 @@ +import Mathlib.Tactic.Positivity import Mathlib.Data.Complex.Exponential import Mathlib.Data.Real.Sqrt import Mathlib.Analysis.Normed.Group.Basic import Mathlib.Analysis.SpecialFunctions.Pow.Real import Mathlib.Analysis.SpecialFunctions.Log.Basic import Mathlib.MeasureTheory.Integral.Bochner -import Mathlib.Tactic.Positivity.Finset /-! # Tests for the `positivity` tactic From 21add2055a9b9a9308ecf80d084d273c8757b699 Mon Sep 17 00:00:00 2001 From: damiano Date: Thu, 21 Nov 2024 03:15:49 +0000 Subject: [PATCH 168/829] fix: silence the header linter on non-`Mathlib.lean` imported files (#19267) [Zulip](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/intrusive.20linters/near/483368456) --- Mathlib/Tactic/Linter/Header.lean | 5 ++++- 1 file changed, 4 insertions(+), 1 deletion(-) diff --git a/Mathlib/Tactic/Linter/Header.lean b/Mathlib/Tactic/Linter/Header.lean index 3e2a44a3d836f..9f5be314fe2ff 100644 --- a/Mathlib/Tactic/Linter/Header.lean +++ b/Mathlib/Tactic/Linter/Header.lean @@ -290,7 +290,10 @@ def duplicateImportsCheck (imports : Array Syntax) : CommandElabM Unit := do @[inherit_doc Mathlib.Linter.linter.style.header] def headerLinter : Linter where run := withSetOptionIn fun stx ↦ do let mainModule ← getMainModule - unless Linter.getLinterValue linter.style.header (← getOptions) || (← isInMathlib mainModule) do + -- The linter skips files not imported in `Mathlib.lean`, to avoid linting "scratch files". + -- However, it is active in the test file `MathlibTest.Header` for the linter itself. + unless (← isInMathlib mainModule) || mainModule == `MathlibTest.Header do return + unless Linter.getLinterValue linter.style.header (← getOptions) do return if (← get).messages.hasErrors then return From ebe784ec742f8a4ad2f6e96c6084ceadaeff659a Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Thu, 21 Nov 2024 03:58:26 +0000 Subject: [PATCH 169/829] fix: more stable choice of representative for atoms in `ring` and `abel` (#19119) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Algebraic normalization tactics (`ring`, `abel`, etc.) typically require a concept of "atom", with expressions which are t-defeq (for some transparency t) being identified. However, the particular representative of this equivalence class which turns up in the final normalized expression is basically random. (It often ends up being the "first", in some tree sense, occurrence of the equivalence class in the expression being normalized.) This ends up being particularly unpredictable when multiple expressions are being normalized simultaneously, e.g. with `ring_nf` or `abel_nf`: it can occur that in different expressions, a different representative of the equivalence class is chosen. For example, on current Mathlib, ```lean example (x : ℤ) (R : ℤ → ℤ → Prop) : True := by let a := x have h : R (a + x) (x + a) := sorry ring_nf at h ``` the statement of `h` after the ring-normalization step is `h : R (a * 2) (x * 2)`. Here `a` and `x` are reducibly defeq. When normalizing `a + x`, the representative `a` was chosen for the equivalence class; when normalizing `x + a`, the representative `x` was chosen. This PR implements a fix. The `AtomM` monad (which is used for atom-tracking in `ring`, `abel`, etc.) has its `addAtom` function modified to report, not just the index of an atom, but also the stored form of the atom. Then the code surrounding `addAtom` calls in the `ring`, `abel` and `module` tactics is modified to use the stored form of the atom, rather than the form actually encountered at that point in the expression. Co-authored-by: Eric Wieser --- Mathlib/Tactic/Abel.lean | 6 ++--- Mathlib/Tactic/ITauto.lean | 2 +- Mathlib/Tactic/Linarith/Preprocessing.lean | 6 ++--- Mathlib/Tactic/Module.lean | 6 ++--- Mathlib/Tactic/Polyrith.lean | 2 +- Mathlib/Tactic/Ring/Basic.lean | 16 +++++------ Mathlib/Tactic/TFAE.lean | 10 +++---- Mathlib/Util/AtomM.lean | 31 ++++++++++++++++++---- MathlibTest/abel.lean | 29 ++++++++++++++++++++ MathlibTest/ring.lean | 25 +++++++++++++++++ 10 files changed, 103 insertions(+), 30 deletions(-) diff --git a/Mathlib/Tactic/Abel.lean b/Mathlib/Tactic/Abel.lean index dafcc98e2ffe0..4002862603dbc 100644 --- a/Mathlib/Tactic/Abel.lean +++ b/Mathlib/Tactic/Abel.lean @@ -247,11 +247,11 @@ theorem term_atom_pfg {α} [AddCommGroup α] (x x' : α) (h : x = x') : x = term /-- Interpret an expression as an atom for `abel`'s normal form. -/ def evalAtom (e : Expr) : M (NormalExpr × Expr) := do let { expr := e', proof?, .. } ← (← readThe AtomM.Context).evalAtom e - let i ← AtomM.addAtom e' + let (i, a) ← AtomM.addAtom e' let p ← match proof? with | none => iapp ``term_atom #[e] - | some p => iapp ``term_atom_pf #[e, e', p] - return (← term' (← intToExpr 1, 1) (i, e') (← zero'), p) + | some p => iapp ``term_atom_pf #[e, a, p] + return (← term' (← intToExpr 1, 1) (i, a) (← zero'), p) theorem unfold_sub {α} [SubtractionMonoid α] (a b c : α) (h : a + -b = c) : a - b = c := by rw [sub_eq_add_neg, h] diff --git a/Mathlib/Tactic/ITauto.lean b/Mathlib/Tactic/ITauto.lean index f0449d9828b14..f87cbcddd6ec3 100644 --- a/Mathlib/Tactic/ITauto.lean +++ b/Mathlib/Tactic/ITauto.lean @@ -486,7 +486,7 @@ partial def reify (e : Q(Prop)) : AtomM IProp := | ~q(@Ne Prop $a $b) => return .not (.eq (← reify a) (← reify b)) | e => if e.isArrow then return .imp (← reify e.bindingDomain!) (← reify e.bindingBody!) - else return .var (← AtomM.addAtom e) + else return .var (← AtomM.addAtom e).1 /-- Once we have a proof object, we have to apply it to the goal. -/ partial def applyProof (g : MVarId) (Γ : NameMap Expr) (p : Proof) : MetaM Unit := diff --git a/Mathlib/Tactic/Linarith/Preprocessing.lean b/Mathlib/Tactic/Linarith/Preprocessing.lean index eca9903c6a74d..eda221c8c33d1 100644 --- a/Mathlib/Tactic/Linarith/Preprocessing.lean +++ b/Mathlib/Tactic/Linarith/Preprocessing.lean @@ -274,12 +274,12 @@ partial def findSquares (s : RBSet (Nat × Bool) lexOrd.compare) (e : Expr) : | (``HPow.hPow, #[_, _, _, _, a, b]) => match b.numeral? with | some 2 => do let s ← findSquares s a - let ai ← AtomM.addAtom a + let (ai, _) ← AtomM.addAtom a return (s.insert (ai, true)) | _ => e.foldlM findSquares s | (``HMul.hMul, #[_, _, _, _, a, b]) => do - let ai ← AtomM.addAtom a - let bi ← AtomM.addAtom b + let (ai, _) ← AtomM.addAtom a + let (bi, _) ← AtomM.addAtom b if ai = bi then do let s ← findSquares s a return (s.insert (ai, false)) diff --git a/Mathlib/Tactic/Module.lean b/Mathlib/Tactic/Module.lean index 4df4c398c9ef4..e06d5ce2088c2 100644 --- a/Mathlib/Tactic/Module.lean +++ b/Mathlib/Tactic/Module.lean @@ -464,9 +464,9 @@ partial def parse (iM : Q(AddCommMonoid $M)) (x : Q($M)) : pure ⟨0, q(Nat), q(Nat.instSemiring), q(AddCommGroup.toNatModule), [], q(NF.zero_eq_eval $M)⟩ /- anything else should be treated as an atom -/ | _ => - let k : ℕ ← AtomM.addAtom x - pure ⟨0, q(Nat), q(Nat.instSemiring), q(AddCommGroup.toNatModule), [((q(1), x), k)], - q(NF.atom_eq_eval $x)⟩ + let (k, ⟨x', _⟩) ← AtomM.addAtomQ x + pure ⟨0, q(Nat), q(Nat.instSemiring), q(AddCommGroup.toNatModule), [((q(1), x'), k)], + q(NF.atom_eq_eval $x')⟩ /-- Given expressions `R` and `M` representing types such that `M`'s is a module over `R`'s, and given two terms `l₁`, `l₂` of type `qNF R M`, i.e. lists of `(Q($R) × Q($M)) × ℕ`s (two `Expr`s diff --git a/Mathlib/Tactic/Polyrith.lean b/Mathlib/Tactic/Polyrith.lean index 57ef95ad04cbb..50e42c8eecbd1 100644 --- a/Mathlib/Tactic/Polyrith.lean +++ b/Mathlib/Tactic/Polyrith.lean @@ -133,7 +133,7 @@ partial def parse {u : Level} {α : Q(Type u)} (sα : Q(CommSemiring $α)) (c : Ring.Cache sα) (e : Q($α)) : AtomM Poly := do let els := do try pure <| Poly.const (← (← NormNum.derive e).toRat) - catch _ => pure <| Poly.var (← addAtom e) + catch _ => pure <| Poly.var (← addAtom e).1 let .const n _ := (← withReducible <| whnf e).getAppFn | els match n, c.rα with | ``HAdd.hAdd, _ | ``Add.add, _ => match e with diff --git a/Mathlib/Tactic/Ring/Basic.lean b/Mathlib/Tactic/Ring/Basic.lean index 84faee3f1a0d4..b5dacb13e9ab5 100644 --- a/Mathlib/Tactic/Ring/Basic.lean +++ b/Mathlib/Tactic/Ring/Basic.lean @@ -512,9 +512,8 @@ mutual partial def ExBase.evalNatCast {a : Q(ℕ)} (va : ExBase sℕ a) : AtomM (Result (ExBase sα) q($a)) := match va with | .atom _ => do - let a' : Q($α) := q($a) - let i ← addAtom a' - pure ⟨a', ExBase.atom i, (q(Eq.refl $a') : Expr)⟩ + let (i, ⟨b', _⟩) ← addAtomQ q($a) + pure ⟨b', ExBase.atom i, q(Eq.refl $b')⟩ | .sum va => do let ⟨_, vc, p⟩ ← va.evalNatCast pure ⟨_, .sum vc, p⟩ @@ -952,11 +951,11 @@ Evaluates an atom, an expression where `ring` can find no additional structure. def evalAtom (e : Q($α)) : AtomM (Result (ExSum sα) e) := do let r ← (← read).evalAtom e have e' : Q($α) := r.expr - let i ← addAtom e' - let ve' := (ExBase.atom i (e := e')).toProd (ExProd.mkNat sℕ 1).2 |>.toSum + let (i, ⟨a', _⟩) ← addAtomQ e' + let ve' := (ExBase.atom i (e := a')).toProd (ExProd.mkNat sℕ 1).2 |>.toSum pure ⟨_, ve', match r.proof? with | none => (q(atom_pf $e) : Expr) - | some (p : Q($e = $e')) => (q(atom_pf' $p) : Expr)⟩ + | some (p : Q($e = $a')) => (q(atom_pf' $p) : Expr)⟩ theorem inv_mul {R} [DivisionRing R] {a₁ a₂ a₃ b₁ b₃ c} (_ : (a₁⁻¹ : R) = b₁) (_ : (a₃⁻¹ : R) = b₃) @@ -977,9 +976,8 @@ variable (dα : Q(DivisionRing $α)) /-- Applies `⁻¹` to a polynomial to get an atom. -/ def evalInvAtom (a : Q($α)) : AtomM (Result (ExBase sα) q($a⁻¹)) := do - let a' : Q($α) := q($a⁻¹) - let i ← addAtom a' - pure ⟨a', ExBase.atom i, (q(Eq.refl $a') : Expr)⟩ + let (i, ⟨b', _⟩) ← addAtomQ q($a⁻¹) + pure ⟨b', ExBase.atom i, q(Eq.refl $b')⟩ /-- Inverts a polynomial `va` to get a normalized result polynomial. diff --git a/Mathlib/Tactic/TFAE.lean b/Mathlib/Tactic/TFAE.lean index 45405bfe5ad58..ac5f1114a8529 100644 --- a/Mathlib/Tactic/TFAE.lean +++ b/Mathlib/Tactic/TFAE.lean @@ -298,18 +298,18 @@ elab_rules : tactic goal.withContext do let (tfaeListQ, tfaeList) ← getTFAEList (← goal.getType) closeMainGoal `tfae_finish <|← AtomM.run .reducible do - let is ← tfaeList.mapM AtomM.addAtom + let is ← tfaeList.mapM (fun e ↦ Prod.fst <$> AtomM.addAtom e) let mut hyps := #[] for hyp in ← getLocalHyps do let ty ← whnfR <|← instantiateMVars <|← inferType hyp if let (``Iff, #[p1, p2]) := ty.getAppFnArgs then - let q1 ← AtomM.addAtom p1 - let q2 ← AtomM.addAtom p2 + let (q1, _) ← AtomM.addAtom p1 + let (q2, _) ← AtomM.addAtom p2 hyps := hyps.push (q1, q2, ← mkAppM ``Iff.mp #[hyp]) hyps := hyps.push (q2, q1, ← mkAppM ``Iff.mpr #[hyp]) else if ty.isArrow then - let q1 ← AtomM.addAtom ty.bindingDomain! - let q2 ← AtomM.addAtom ty.bindingBody! + let (q1, _) ← AtomM.addAtom ty.bindingDomain! + let (q2, _) ← AtomM.addAtom ty.bindingBody! hyps := hyps.push (q1, q2, hyp) proveTFAE hyps (← get).atoms is tfaeListQ diff --git a/Mathlib/Util/AtomM.lean b/Mathlib/Util/AtomM.lean index 0b829028c0b11..bc85531a753a4 100644 --- a/Mathlib/Util/AtomM.lean +++ b/Mathlib/Util/AtomM.lean @@ -5,6 +5,7 @@ Authors: Mario Carneiro -/ import Mathlib.Init import Lean.Meta.Tactic.Simp.Types +import Qq /-! # A monad for tracking and deduplicating atoms @@ -44,13 +45,33 @@ def AtomM.run {α : Type} (red : TransparencyMode) (m : AtomM α) MetaM α := (m { red, evalAtom }).run' {} -/-- Get the index corresponding to an atomic expression, if it has already been encountered, or -put it in the list of atoms and return the new index, otherwise. -/ -def AtomM.addAtom (e : Expr) : AtomM Nat := do +/-- If an atomic expression has already been encountered, get the index and the stored form of the +atom (which will be defeq at the specified transparency, but not necessarily syntactically equal). +If the atomic expression has *not* already been encountered, store it in the list of atoms, and +return the new index (and the stored form of the atom, which will be itself). + +In a normalizing tactic, the expression returned by `addAtom` should be considered the normal form. +-/ +def AtomM.addAtom (e : Expr) : AtomM (Nat × Expr) := do let c ← get for h : i in [:c.atoms.size] do if ← withTransparency (← read).red <| isDefEq e c.atoms[i] then - return i - modifyGet fun c ↦ (c.atoms.size, { c with atoms := c.atoms.push e }) + return (i, c.atoms[i]) + modifyGet fun c ↦ ((c.atoms.size, e), { c with atoms := c.atoms.push e }) + +open Qq in +/-- If an atomic expression has already been encountered, get the index and the stored form of the +atom (which will be defeq at the specified transparency, but not necessarily syntactically equal). +If the atomic expression has *not* already been encountered, store it in the list of atoms, and +return the new index (and the stored form of the atom, which will be itself). + +In a normalizing tactic, the expression returned by `addAtomQ` should be considered the normal form. + +This is a strongly-typed version of `AtomM.addAtom` for code using `Qq`. +-/ +def AtomM.addAtomQ {u : Level} {α : Q(Type u)} (e : Q($α)) : + AtomM (Nat × {e' : Q($α) // $e =Q $e'}) := do + let (n, e') ← AtomM.addAtom e + return (n, ⟨e', ⟨⟩⟩) end Mathlib.Tactic diff --git a/MathlibTest/abel.lean b/MathlibTest/abel.lean index d59304f17cbfb..f2235b46071cf 100644 --- a/MathlibTest/abel.lean +++ b/MathlibTest/abel.lean @@ -141,3 +141,32 @@ example [AddCommGroup α] (x y z : α) (h : False) (w : x - x = y + z) : False : abel_nf at * guard_hyp w : 0 = y + z assumption + +section +abbrev myId (a : ℤ) : ℤ := a + +/- +Test that when `abel_nf` normalizes multiple expressions which contain a particular atom, it uses a +form for that atom which is consistent between expressions. + +We can't use `guard_hyp h :ₛ` here, as while it does tell apart `x` and `myId x`, it also complains +about differing instance paths. +-/ +/-- +info: α : Type _ +a b : α +x : ℤ +R : ℤ → ℤ → Prop +hR : Reflexive R +h : R (2 • myId x) (2 • myId x) +⊢ True +-/ +#guard_msgs (info) in +set_option pp.mvars false in +example (x : ℤ) (R : ℤ → ℤ → Prop) (hR : Reflexive R) : True := by + have h : R (myId x + x) (x + myId x) := hR .. + abel_nf at h + trace_state + trivial + +end diff --git a/MathlibTest/ring.lean b/MathlibTest/ring.lean index ffdfe755e4dec..8a0a07fea0178 100644 --- a/MathlibTest/ring.lean +++ b/MathlibTest/ring.lean @@ -178,3 +178,28 @@ example {n : ℝ} (_hn : 0 ≤ n) : (n + 1 / 2) ^ 2 * (n + 1 + 1 / 3) ≤ (n + 1 ring_nf trace_state exact test_sorry + +section +abbrev myId (a : ℤ) : ℤ := a + +/- +Test that when `ring_nf` normalizes multiple expressions which contain a particular atom, it uses a +form for that atom which is consistent between expressions. + +We can't use `guard_hyp h :ₛ` here, as while it does tell apart `x` and `myId x`, it also complains +about differing instance paths. +-/ +/-- +info: x : ℤ +R : ℤ → ℤ → Prop +h : R (myId x * 2) (myId x * 2) +⊢ True +-/ +#guard_msgs (info) in +example (x : ℤ) (R : ℤ → ℤ → Prop) : True := by + have h : R (myId x + x) (x + myId x) := test_sorry + ring_nf at h + trace_state + trivial + +end From bc56b5cdfe035473d40355c7f30a3f837e98db9a Mon Sep 17 00:00:00 2001 From: damiano Date: Thu, 21 Nov 2024 04:49:01 +0000 Subject: [PATCH 170/829] chore(CI): remove redundant `bors x` actions, clean up new one (#19313) I did a fair amount of testing, making sure that the new action has the same functionality as the old one. As far as I can tell, they work the same. The only exception could be that the new action does not remove the `awaiting-author` label, since all the tests that I saw did not have the label in the first place. --- .github/workflows/add_label_from_comment.yml | 60 ------------------ .github/workflows/add_label_from_review.yml | 62 ------------------- .../add_label_from_review_comment.yml | 62 ------------------- .github/workflows/maintainer_bors.yml | 16 +---- 4 files changed, 3 insertions(+), 197 deletions(-) delete mode 100644 .github/workflows/add_label_from_comment.yml delete mode 100644 .github/workflows/add_label_from_review.yml delete mode 100644 .github/workflows/add_label_from_review_comment.yml diff --git a/.github/workflows/add_label_from_comment.yml b/.github/workflows/add_label_from_comment.yml deleted file mode 100644 index 30a161785134f..0000000000000 --- a/.github/workflows/add_label_from_comment.yml +++ /dev/null @@ -1,60 +0,0 @@ -name: Add "ready-to-merge" and "delegated" label from comment - -on: - issue_comment: - types: [created] - -jobs: - add_ready_to_merge_label: - name: Add ready-to-merge label - if: github.event.issue.pull_request && (startsWith(github.event.comment.body, 'bors r+') || contains(toJSON(github.event.comment.body), '\nbors r+') || startsWith(github.event.comment.body, 'bors merge') || contains(toJSON(github.event.comment.body), '\nbors merge')) - runs-on: ubuntu-latest - steps: - - id: user_permission - uses: actions-cool/check-user-permission@v2 - with: - require: 'admin' - - - if: (steps.user_permission.outputs.require-result == 'true') || (github.event.comment.user.login == 'leanprover-community-mathlib4-bot') || (github.event.comment.user.login == 'leanprover-community-bot-assistant') - uses: octokit/request-action@v2.x - id: add_label - name: Add label - with: - route: POST /repos/:repository/issues/:issue_number/labels - repository: ${{ github.repository }} - issue_number: ${{ github.event.issue.number }} - labels: '["ready-to-merge"]' - env: - GITHUB_TOKEN: ${{ secrets.GITHUB_TOKEN }} - - - if: (steps.user_permission.outputs.require-result == 'true') || (github.event.comment.user.login == 'leanprover-community-mathlib4-bot') || (github.event.comment.user.login == 'leanprover-community-bot-assistant') - id: remove_labels - name: Remove "awaiting-author" - # we use curl rather than octokit/request-action so that the job won't fail - # (and send an annoying email) if the labels don't exist - run: | - curl --request DELETE \ - --url https://api.github.com/repos/${{ github.repository }}/issues/${{ github.event.issue.number }}/labels/awaiting-author \ - --header 'authorization: Bearer ${{ secrets.GITHUB_TOKEN }}' - - add_delegated_label: - name: Add delegated label - if: github.event.issue.pull_request && (startsWith(github.event.comment.body, 'bors d') || contains(toJSON(github.event.comment.body), '\nbors d')) - runs-on: ubuntu-latest - steps: - - id: user_permission - uses: actions-cool/check-user-permission@v2 - with: - require: 'admin' - - - if: (steps.user_permission.outputs.require-result == 'true') || (github.event.comment.user.login == 'leanprover-community-mathlib4-bot') || (github.event.comment.user.login == 'leanprover-community-bot-assistant') - uses: octokit/request-action@v2.x - id: add_label - name: Add label - with: - route: POST /repos/:repository/issues/:issue_number/labels - repository: ${{ github.repository }} - issue_number: ${{ github.event.issue.number }} - labels: '["delegated"]' - env: - GITHUB_TOKEN: ${{ secrets.GITHUB_TOKEN }} diff --git a/.github/workflows/add_label_from_review.yml b/.github/workflows/add_label_from_review.yml deleted file mode 100644 index 9406866ff214e..0000000000000 --- a/.github/workflows/add_label_from_review.yml +++ /dev/null @@ -1,62 +0,0 @@ -name: Add "ready-to-merge" and "delegated" label from PR review - -on: - pull_request_review: - types: [submitted] - -jobs: - add_ready_to_merge_label: - name: Add ready-to-merge label - if: (startsWith(github.event.review.body, 'bors r+') || contains(toJSON(github.event.review.body), '\nbors r+') || startsWith(github.event.review.body, 'bors merge') || contains(toJSON(github.event.review.body), '\nbors merge')) - runs-on: ubuntu-latest - steps: - - id: user_permission - uses: actions-cool/check-user-permission@v2 - with: - require: 'write' - token: ${{ secrets.TRIAGE_TOKEN }} - - - if: (steps.user_permission.outputs.require-result == 'true') || (github.event.review.user.login == 'leanprover-community-mathlib4-bot') || (github.event.review.user.login == 'leanprover-community-bot-assistant') - uses: octokit/request-action@v2.x - id: add_label - name: Add label - with: - route: POST /repos/:repository/issues/:issue_number/labels - repository: ${{ github.repository }} - issue_number: ${{ github.event.pull_request.number }} - labels: '["ready-to-merge"]' - env: - GITHUB_TOKEN: ${{ secrets.TRIAGE_TOKEN }} - - - if: (steps.user_permission.outputs.require-result == 'true') || (github.event.review.user.login == 'leanprover-community-mathlib4-bot') || (github.event.review.user.login == 'leanprover-community-bot-assistant') - id: remove_labels - name: Remove "awaiting-author" - # we use curl rather than octokit/request-action so that the job won't fail - # (and send an annoying email) if the labels don't exist - run: | - curl --request DELETE \ - --url https://api.github.com/repos/${{ github.repository }}/issues/${{ github.event.pull_request.number }}/labels/awaiting-author \ - --header 'authorization: Bearer ${{ secrets.TRIAGE_TOKEN }}' - - add_delegated_label: - name: Add delegated label - if: (startsWith(github.event.review.body, 'bors d') || contains(toJSON(github.event.review.body), '\nbors d')) - runs-on: ubuntu-latest - steps: - - id: user_permission - uses: actions-cool/check-user-permission@v2 - with: - require: 'write' - token: ${{ secrets.TRIAGE_TOKEN }} - - - if: (steps.user_permission.outputs.require-result == 'true') || (github.event.review.user.login == 'leanprover-community-mathlib4-bot') || (github.event.review.user.login == 'leanprover-community-bot-assistant') - uses: octokit/request-action@v2.x - id: add_label - name: Add label - with: - route: POST /repos/:repository/issues/:issue_number/labels - repository: ${{ github.repository }} - issue_number: ${{ github.event.pull_request.number }} - labels: '["delegated"]' - env: - GITHUB_TOKEN: ${{ secrets.TRIAGE_TOKEN }} diff --git a/.github/workflows/add_label_from_review_comment.yml b/.github/workflows/add_label_from_review_comment.yml deleted file mode 100644 index da8b643fb0ee4..0000000000000 --- a/.github/workflows/add_label_from_review_comment.yml +++ /dev/null @@ -1,62 +0,0 @@ -name: Add "ready-to-merge" and "delegated" label from PR review comment - -on: - pull_request_review_comment: - types: [created] - -jobs: - add_ready_to_merge_label: - name: Add ready-to-merge label - if: (startsWith(github.event.comment.body, 'bors r+') || contains(toJSON(github.event.comment.body), '\nbors r+') || startsWith(github.event.comment.body, 'bors merge') || contains(toJSON(github.event.comment.body), '\nbors merge')) - runs-on: ubuntu-latest - steps: - - id: user_permission - uses: actions-cool/check-user-permission@v2 - with: - require: 'write' - token: ${{ secrets.TRIAGE_TOKEN }} - - - if: (steps.user_permission.outputs.require-result == 'true') || (github.event.comment.user.login == 'leanprover-community-mathlib4-bot') || (github.event.comment.user.login == 'leanprover-community-bot-assistant') - uses: octokit/request-action@v2.x - id: add_label - name: Add label - with: - route: POST /repos/:repository/issues/:issue_number/labels - repository: ${{ github.repository }} - issue_number: ${{ github.event.pull_request.number }} - labels: '["ready-to-merge"]' - env: - GITHUB_TOKEN: ${{ secrets.TRIAGE_TOKEN }} - - - if: (steps.user_permission.outputs.require-result == 'true') || (github.event.comment.user.login == 'leanprover-community-mathlib4-bot') || (github.event.comment.user.login == 'leanprover-community-bot-assistant') - id: remove_labels - name: Remove "awaiting-author" - # we use curl rather than octokit/request-action so that the job won't fail - # (and send an annoying email) if the labels don't exist - run: | - curl --request DELETE \ - --url https://api.github.com/repos/${{ github.repository }}/issues/${{ github.event.pull_request.number }}/labels/awaiting-author \ - --header 'authorization: Bearer ${{ secrets.TRIAGE_TOKEN }}' - - add_delegated_label: - name: Add delegated label - if: (startsWith(github.event.comment.body, 'bors d') || contains(toJSON(github.event.comment.body), '\nbors d')) - runs-on: ubuntu-latest - steps: - - id: user_permission - uses: actions-cool/check-user-permission@v2 - with: - require: 'write' - token: ${{ secrets.TRIAGE_TOKEN }} - - - if: (steps.user_permission.outputs.require-result == 'true') || (github.event.comment.user.login == 'leanprover-community-mathlib4-bot') || (github.event.comment.user.login == 'leanprover-community-bot-assistant') - uses: octokit/request-action@v2.x - id: add_label - name: Add label - with: - route: POST /repos/:repository/issues/:issue_number/labels - repository: ${{ github.repository }} - issue_number: ${{ github.event.pull_request.number }} - labels: '["delegated"]' - env: - GITHUB_TOKEN: ${{ secrets.TRIAGE_TOKEN }} diff --git a/.github/workflows/maintainer_bors.yml b/.github/workflows/maintainer_bors.yml index 073477b4fa324..8b06a44a1b926 100644 --- a/.github/workflows/maintainer_bors.yml +++ b/.github/workflows/maintainer_bors.yml @@ -19,13 +19,8 @@ jobs: # both set simultaneously: depending on the event that triggers the PR, usually only one is set env: AUTHOR: ${{ github.event.comment.user.login }}${{ github.event.review.user.login }} - PR_NUMBER: ${{ github.event.issue.number }}${{ github.event.pull_request.number }} COMMENT_EVENT: ${{ github.event.comment.body }} COMMENT_REVIEW: ${{ github.event.review.body }} - PR_TITLE_ISSUE: ${{ github.event.issue.title }} - PR_TITLE_PR: ${{ github.event.pull_request.title }} - PR_URL: ${{ github.event.issue.html_url }}${{ github.event.pull_request.html_url }} - EVENT_NAME: ${{ github.event_name }} name: Add ready-to-merge or delegated label runs-on: ubuntu-latest steps: @@ -43,8 +38,7 @@ jobs: printf $'"bors delegate" or "bors merge" found? \'%s\'\n' "${m_or_d}" printf $'AUTHOR: \'%s\'\n' "${AUTHOR}" - printf $'PR_NUMBER: \'%s\'\n' "${PR_NUMBER}" - printf $'OTHER_NUMBER: \'%s\'\n' "${{ github.event.issue.number }}${{ github.event.pull_request.number }}" + printf $'PR_NUMBER: \'%s\'\n' "${{ github.event.issue.number }}${{ github.event.pull_request.number }}" printf $'%s' "${{ github.event.issue.number }}${{ github.event.pull_request.number }}" | hexdump -cC printf $'mOrD=%s\n' "${m_or_d}" >> "${GITHUB_OUTPUT}" @@ -64,9 +58,6 @@ jobs: uses: actions-cool/check-user-permission@v2 with: require: 'admin' - # review(_comment) use - # require: 'write' - # token: ${{ secrets.TRIAGE_TOKEN }} - name: Add ready-to-merge or delegated label id: add_label @@ -75,13 +66,12 @@ jobs: steps.merge_or_delegate.outputs.bot == 'true' ) }} uses: octokit/request-action@v2.x with: - # check is this ok? was /repos/:repository/issues/:issue_number/labels route: POST /repos/:repository/issues/:issue_number/labels repository: ${{ github.repository }} issue_number: ${{ github.event.issue.number }}${{ github.event.pull_request.number }} labels: '["${{ steps.merge_or_delegate.outputs.mOrD }}"]' env: - GITHUB_TOKEN: ${{ secrets.TRIAGE_TOKEN }} # comment uses ${{ secrets.GITHUB_TOKEN }} + GITHUB_TOKEN: ${{ secrets.TRIAGE_TOKEN }} - if: ${{ ! steps.merge_or_delegate.outputs.mOrD == '' && ( steps.user_permission.outputs.require-result == 'true' || @@ -92,6 +82,6 @@ jobs: # (and send an annoying email) if the labels don't exist run: | curl --request DELETE \ - --url "https://api.github.com/repos/${{ github.repository }}/issues/${PR_NUMBER}/labels/awaiting-author" \ + --url "https://api.github.com/repos/${{ github.repository }}/issues/${{ github.event.issue.number }}${{ github.event.pull_request.number }}/labels/awaiting-author" \ --header 'authorization: Bearer ${{ secrets.TRIAGE_TOKEN }}' # comment uses ${{ secrets.GITHUB_TOKEN }} From 79ca1675422026fe7077a20433f658a9ef959078 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Thu, 21 Nov 2024 05:46:14 +0000 Subject: [PATCH 171/829] chore(discover-lean-pr-testing): determine old and new lean-toolchain more accurately (#19316) Also, improve the way we fetch the history between the old and new nightlies. --- .../workflows/discover-lean-pr-testing.yml | 32 +++++++++++++++---- 1 file changed, 25 insertions(+), 7 deletions(-) diff --git a/.github/workflows/discover-lean-pr-testing.yml b/.github/workflows/discover-lean-pr-testing.yml index 3c414d9e99689..69e46265bbbc0 100644 --- a/.github/workflows/discover-lean-pr-testing.yml +++ b/.github/workflows/discover-lean-pr-testing.yml @@ -24,6 +24,23 @@ jobs: git config --global user.name "github-actions" git config --global user.email "github-actions@github.com" + - name: Determine old and new lean-toolchain + id: determine-toolchains + run: | + # `lean-toolchain` contains a string of the form "leanprover/lean4:nightly-2024-11-20" + # We grab the part after the ":" + NEW=$(cut -f2 -d: lean-toolchain) + + # Find the commit hash of the previous change to `lean-toolchain` + PREV_COMMIT=$(git log -2 --format=format:%H -- lean-toolchain | tail -1) + + # Get the contents of `lean-toolchain` from the previous commit + # The "./" in front of the path is important for `git show` + OLD=$(git show "$PREV_COMMIT":./lean-toolchain | cut -f2 -d:) + + echo "new=$NEW" >> "$GITHUB_OUTPUT" + echo "old=$OLD" >> "$GITHUB_OUTPUT" + - name: Clone lean4 repository and get PRs id: get-prs run: | @@ -38,14 +55,15 @@ jobs: # Navigate to the cloned repository cd lean4-nightly || exit 1 - # Shallow fetch of the last two tags - LAST_TWO_TAGS=$(git ls-remote --tags | grep nightly | tail -2 | cut -f2) - OLD=$(echo "$LAST_TWO_TAGS" | head -1) - NEW=$(echo "$LAST_TWO_TAGS" | tail -1) + # Use the OLD and NEW toolchains determined in the previous step + OLD="${{ steps.determine-toolchains.outputs.old }}" + NEW="${{ steps.determine-toolchains.outputs.new }}" + + # Fetch the $OLD tag git fetch --depth=1 origin tag "$OLD" --no-tags - # Fetch the latest 100 commits prior to the $NEW tag. - # This should cover the terrain between $OLD and $NEW while still being a speedy download. - git fetch --depth=100 origin tag "$NEW" --no-tags + # Fetch the $NEW tag. + # This will only fetch commits newer than previously fetched commits (ie $OLD) + git fetch origin tag "$NEW" --no-tags PRS=$(git log --oneline "$OLD..$NEW" | sed 's/.*(#\([0-9]\+\))$/\1/') From 2a1fef1fb7d545e232fd5476d36a88f2f22afa0e Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Thu, 21 Nov 2024 05:54:56 +0000 Subject: [PATCH 172/829] chore: update Mathlib dependencies 2024-11-21 (#19317) 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 d40bdeb3be7db..8fc2bedf751bd 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "485efbc439ee0ebdeae8afb0acd24a5e82e2f771", + "rev": "a822446d61ad7e7f5e843365c7041c326553050a", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From ab49b8360a2fb3a08e63cf0c0bc344d4827eeb53 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Thu, 21 Nov 2024 06:49:22 +0000 Subject: [PATCH 173/829] chore: assorted golf (#19319) --- Archive/Imo/Imo2008Q2.lean | 7 +------ Archive/Imo/Imo2013Q5.lean | 12 +++--------- Mathlib/Analysis/Calculus/Monotone.lean | 10 +++++----- Mathlib/Analysis/Convex/Uniform.lean | 11 +---------- Mathlib/Analysis/PSeries.lean | 7 ++----- .../SpecialFunctions/Gamma/BohrMollerup.lean | 2 +- Mathlib/Analysis/SpecialFunctions/Log/Monotone.lean | 12 +++++------- Mathlib/NumberTheory/FermatPsp.lean | 6 ++---- Mathlib/NumberTheory/Modular.lean | 7 ++----- Mathlib/Topology/MetricSpace/Kuratowski.lean | 3 +-- 10 files changed, 23 insertions(+), 54 deletions(-) diff --git a/Archive/Imo/Imo2008Q2.lean b/Archive/Imo/Imo2008Q2.lean index 0340c3f70479b..41687a4a977a0 100644 --- a/Archive/Imo/Imo2008Q2.lean +++ b/Archive/Imo/Imo2008Q2.lean @@ -106,12 +106,7 @@ theorem imo2008_q2b : Set.Infinite rationalSolutions := by exact ⟨rfl, rfl, rfl⟩ · have hg : -z = g (x, y, z) := rfl rw [hg, hz_def]; ring - have h₂ : q < t * (t + 1) := by - calc - q < q + 1 := by linarith - _ ≤ t := le_max_left (q + 1) 1 - _ ≤ t + t ^ 2 := by linarith [sq_nonneg t] - _ = t * (t + 1) := by ring + have h₂ : q < t * (t + 1) := by linarith [sq_nonneg t, le_max_left (q + 1) 1] exact ⟨h₁, h₂⟩ have hK_inf : Set.Infinite K := by intro h; apply hK_not_bdd; exact Set.Finite.bddAbove h exact hK_inf.of_image g diff --git a/Archive/Imo/Imo2013Q5.lean b/Archive/Imo/Imo2013Q5.lean index 225e64ae35b02..eb996b73633a4 100644 --- a/Archive/Imo/Imo2013Q5.lean +++ b/Archive/Imo/Imo2013Q5.lean @@ -65,15 +65,9 @@ theorem le_of_all_pow_lt_succ' {x y : ℝ} (hx : 1 < x) (hy : 0 < y) refine le_of_all_pow_lt_succ hx ?_ h by_contra! hy'' : y ≤ 1 -- Then there exists y' such that 0 < y ≤ 1 < y' < x. - let y' := (x + 1) / 2 - have h_y'_lt_x : y' < x := - calc - (x + 1) / 2 < x * 2 / 2 := by linarith - _ = x := by field_simp - have h1_lt_y' : 1 < y' := - calc - 1 = 1 * 2 / 2 := by field_simp - _ < (x + 1) / 2 := by linarith + have h_y'_lt_x : (x + 1) / 2 < x := by linarith + have h1_lt_y' : 1 < (x + 1) / 2 := by linarith + set y' := (x + 1) / 2 have h_y_lt_y' : y < y' := by linarith have hh : ∀ n, 0 < n → x ^ n - 1 < y' ^ n := by intro n hn diff --git a/Mathlib/Analysis/Calculus/Monotone.lean b/Mathlib/Analysis/Calculus/Monotone.lean index 599a701b6b211..f54f48fe98117 100644 --- a/Mathlib/Analysis/Calculus/Monotone.lean +++ b/Mathlib/Analysis/Calculus/Monotone.lean @@ -108,7 +108,7 @@ theorem StieltjesFunction.ae_hasDerivAt (f : StieltjesFunction) : filter_upwards [this] rintro y ⟨hy : x - 1 < y, h'y : y < x⟩ rw [mem_Iio] - norm_num; nlinarith + nlinarith -- Deduce the correct limit on the left, by sandwiching. have L4 : Tendsto (fun y => (f y - f x) / (y - x)) (𝓝[<] x) (𝓝 (rnDeriv f.measure volume x).toReal) := by @@ -118,7 +118,7 @@ theorem StieltjesFunction.ae_hasDerivAt (f : StieltjesFunction) : refine div_le_div_of_nonpos_of_le (by linarith) ((sub_le_sub_iff_right _).2 ?_) apply f.mono.le_leftLim have : ↑0 < (x - y) ^ 2 := sq_pos_of_pos (sub_pos.2 hy) - norm_num; linarith + linarith · filter_upwards [self_mem_nhdsWithin] rintro y (hy : y < x) refine div_le_div_of_nonpos_of_le (by linarith) ?_ @@ -161,7 +161,7 @@ theorem Monotone.ae_hasDerivAt {f : ℝ → ℝ} (hf : Monotone f) : filter_upwards [this] rintro y ⟨hy : x < y, h'y : y < x + 1⟩ rw [mem_Ioi] - norm_num; nlinarith + nlinarith -- apply the sandwiching argument, with the helper function and `g` apply tendsto_of_tendsto_of_tendsto_of_le_of_le' this hx.2 · filter_upwards [self_mem_nhdsWithin] with y hy @@ -190,7 +190,7 @@ theorem Monotone.ae_hasDerivAt {f : ℝ → ℝ} (hf : Monotone f) : rintro y hy rw [mem_Ioo] at hy rw [mem_Iio] - norm_num; nlinarith + nlinarith -- apply the sandwiching argument, with `g` and the helper function apply tendsto_of_tendsto_of_tendsto_of_le_of_le' hx.1 this · filter_upwards [self_mem_nhdsWithin] @@ -203,7 +203,7 @@ theorem Monotone.ae_hasDerivAt {f : ℝ → ℝ} (hf : Monotone f) : rw [mem_Iio, ← sub_neg] at hy have : 0 < (y - x) ^ 2 := sq_pos_of_neg hy apply div_le_div_of_nonpos_of_le hy.le - exact (sub_le_sub_iff_right _).2 (hf.rightLim_le (by norm_num; linarith)) + exact (sub_le_sub_iff_right _).2 (hf.rightLim_le (by linarith)) -- conclude global differentiability rw [hasDerivAt_iff_tendsto_slope, slope_fun_def_field, (nhds_left'_sup_nhds_right' x).symm, tendsto_sup] diff --git a/Mathlib/Analysis/Convex/Uniform.lean b/Mathlib/Analysis/Convex/Uniform.lean index cbea3a0fbb080..5c0511976660c 100644 --- a/Mathlib/Analysis/Convex/Uniform.lean +++ b/Mathlib/Analysis/Convex/Uniform.lean @@ -96,16 +96,7 @@ theorem exists_forall_closed_ball_dist_add_le_two_sub (hε : 0 < ε) : _ ≤ 2 - δ + δ' + δ' := (add_le_add_three (h (h₁ _ hx') (h₁ _ hy') hxy') (h₂ _ hx hx'.le) (h₂ _ hy hy'.le)) _ ≤ 2 - δ' := by - dsimp only [δ'] - rw [← le_sub_iff_add_le, ← le_sub_iff_add_le, sub_sub, sub_sub] - refine sub_le_sub_left ?_ _ - ring_nf - rw [← mul_div_cancel₀ δ three_ne_zero] - norm_num - -- Porting note: these three extra lines needed to make `exact` work - have : 3 * (δ / 3) * (1 / 3) = δ / 3 := by linarith - rw [this, mul_comm] - gcongr + suffices δ' ≤ δ / 3 by linarith exact min_le_of_right_le <| min_le_right _ _ theorem exists_forall_closed_ball_dist_add_le_two_mul_sub (hε : 0 < ε) (r : ℝ) : diff --git a/Mathlib/Analysis/PSeries.lean b/Mathlib/Analysis/PSeries.lean index b077709ddf4da..93a65386cf5e2 100644 --- a/Mathlib/Analysis/PSeries.lean +++ b/Mathlib/Analysis/PSeries.lean @@ -386,12 +386,9 @@ theorem sum_Ioc_inv_sq_le_sub {k n : ℕ} (hk : k ≠ 0) (h : k ≤ n) : simp only [sub_eq_add_neg, add_assoc, Nat.cast_add, Nat.cast_one, le_add_neg_iff_add_le, add_le_iff_nonpos_right, neg_add_le_iff_le_add, add_zero] have A : 0 < (n : α) := by simpa using hk.bot_lt.trans_le hn - have B : 0 < (n : α) + 1 := by linarith field_simp - rw [div_le_div_iff₀ _ A, ← sub_nonneg] - · ring_nf - rw [add_comm] - exact B.le + rw [div_le_div_iff₀ _ A] + · linarith · -- Porting note: was `nlinarith` positivity diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/BohrMollerup.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/BohrMollerup.lean index 86a1929d236cc..6a9a40ef2d4ff 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gamma/BohrMollerup.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/BohrMollerup.lean @@ -57,7 +57,7 @@ theorem Gamma_mul_add_mul_le_rpow_Gamma_mul_rpow_Gamma {s t a b : ℝ} (hs : 0 < let f : ℝ → ℝ → ℝ → ℝ := fun c u x => exp (-c * x) * x ^ (c * (u - 1)) have e : IsConjExponent (1 / a) (1 / b) := Real.isConjExponent_one_div ha hb hab have hab' : b = 1 - a := by linarith - have hst : 0 < a * s + b * t := add_pos (mul_pos ha hs) (mul_pos hb ht) + have hst : 0 < a * s + b * t := by positivity -- some properties of f: have posf : ∀ c u x : ℝ, x ∈ Ioi (0 : ℝ) → 0 ≤ f c u x := fun c u x hx => mul_nonneg (exp_pos _).le (rpow_pos_of_pos hx _).le diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Monotone.lean b/Mathlib/Analysis/SpecialFunctions/Log/Monotone.lean index 5aa71529e87e1..c6b70bd4eaffa 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Monotone.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Monotone.lean @@ -53,11 +53,9 @@ theorem log_div_self_rpow_antitoneOn {a : ℝ} (ha : 0 < a) : intro x hex y _ hxy have x_pos : 0 < x := lt_of_lt_of_le (exp_pos (1 / a)) hex have y_pos : 0 < y := by linarith - have x_nonneg : 0 ≤ x := le_trans (le_of_lt (exp_pos (1 / a))) hex - have y_nonneg : 0 ≤ y := by linarith nth_rw 1 [← rpow_one y] nth_rw 1 [← rpow_one x] - rw [← div_self (ne_of_lt ha).symm, div_eq_mul_one_div a a, rpow_mul y_nonneg, rpow_mul x_nonneg, + rw [← div_self (ne_of_lt ha).symm, div_eq_mul_one_div a a, rpow_mul y_pos.le, rpow_mul x_pos.le, log_rpow (rpow_pos_of_pos y_pos a), log_rpow (rpow_pos_of_pos x_pos a), mul_div_assoc, mul_div_assoc, mul_le_mul_left (one_div_pos.mpr ha)] refine log_div_self_antitoneOn ?_ ?_ ?_ @@ -65,14 +63,14 @@ theorem log_div_self_rpow_antitoneOn {a : ℝ} (ha : 0 < a) : convert rpow_le_rpow _ hex (le_of_lt ha) using 1 · rw [← exp_mul] simp only [Real.exp_eq_exp] - field_simp [(ne_of_lt ha).symm] - exact le_of_lt (exp_pos (1 / a)) + field_simp + positivity · simp only [Set.mem_setOf_eq] convert rpow_le_rpow _ (_root_.trans hex hxy) (le_of_lt ha) using 1 · rw [← exp_mul] simp only [Real.exp_eq_exp] - field_simp [(ne_of_lt ha).symm] - exact le_of_lt (exp_pos (1 / a)) + field_simp + positivity gcongr theorem log_div_sqrt_antitoneOn : AntitoneOn (fun x : ℝ => log x / √x) { x | exp 2 ≤ x } := by diff --git a/Mathlib/NumberTheory/FermatPsp.lean b/Mathlib/NumberTheory/FermatPsp.lean index 9fc7ee0822135..b5c6b98a175d6 100644 --- a/Mathlib/NumberTheory/FermatPsp.lean +++ b/Mathlib/NumberTheory/FermatPsp.lean @@ -197,10 +197,8 @@ private theorem psp_from_prime_psp {b : ℕ} (b_ge_two : 2 ≤ b) {p : ℕ} (p_p have hi_p : 1 ≤ p := Nat.one_le_of_lt p_gt_two have hi_bsquared : 0 < b ^ 2 - 1 := by -- Porting note: was `by nlinarith [Nat.one_le_pow 2 b hi_b]` - have h0 := mul_le_mul b_ge_two b_ge_two zero_le_two hi_b.le - have h1 : 1 < 2 * 2 := by omega - have := tsub_pos_of_lt (h1.trans_le h0) - rwa [pow_two] + have := Nat.pow_le_pow_left b_ge_two 2 + omega have hi_bpowtwop : 1 ≤ b ^ (2 * p) := Nat.one_le_pow (2 * p) b hi_b have hi_bpowpsubone : 1 ≤ b ^ (p - 1) := Nat.one_le_pow (p - 1) b hi_b -- Other useful facts diff --git a/Mathlib/NumberTheory/Modular.lean b/Mathlib/NumberTheory/Modular.lean index f4c849d0063d4..ff3eef6b03cda 100644 --- a/Mathlib/NumberTheory/Modular.lean +++ b/Mathlib/NumberTheory/Modular.lean @@ -458,11 +458,8 @@ theorem abs_c_le_one (hz : z ∈ 𝒟ᵒ) (hg : g • z ∈ 𝒟ᵒ) : |g 1 0| specialize this hc linarith intro hc - replace hc : 0 < c ^ 4 := by - change 0 < c ^ (2 * 2); rw [pow_mul]; apply sq_pos_of_pos (sq_pos_of_ne_zero hc) - have h₁ := mul_lt_mul_of_pos_right - (mul_lt_mul'' (three_lt_four_mul_im_sq_of_mem_fdo hg) (three_lt_four_mul_im_sq_of_mem_fdo hz) - (by norm_num) (by norm_num)) hc + have h₁ : 3 * 3 * c ^ 4 < 4 * (g • z).im ^ 2 * (4 * z.im ^ 2) * c ^ 4 := by + gcongr <;> apply three_lt_four_mul_im_sq_of_mem_fdo <;> assumption have h₂ : (c * z.im) ^ 4 / normSq (denom (↑g) z) ^ 2 ≤ 1 := div_le_one_of_le₀ (pow_four_le_pow_two_of_pow_two_le (z.c_mul_im_sq_le_normSq_denom g)) diff --git a/Mathlib/Topology/MetricSpace/Kuratowski.lean b/Mathlib/Topology/MetricSpace/Kuratowski.lean index 5a21609d09845..7b7e6e47ae7f0 100644 --- a/Mathlib/Topology/MetricSpace/Kuratowski.lean +++ b/Mathlib/Topology/MetricSpace/Kuratowski.lean @@ -68,8 +68,7 @@ theorem embeddingOfSubset_isometry (H : DenseRange x) : Isometry (embeddingOfSub apply_rules [add_le_add_left, le_abs_self] _ ≤ 2 * (e / 2) + |embeddingOfSubset x b n - embeddingOfSubset x a n| := by rw [C] - apply_rules [add_le_add, mul_le_mul_of_nonneg_left, hn.le, le_refl] - norm_num + gcongr _ ≤ 2 * (e / 2) + dist (embeddingOfSubset x b) (embeddingOfSubset x a) := by have : |embeddingOfSubset x b n - embeddingOfSubset x a n| ≤ dist (embeddingOfSubset x b) (embeddingOfSubset x a) := by From 1e5f38b9c9889a1375c1007ac707aa325c895a18 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 21 Nov 2024 08:54:19 +0000 Subject: [PATCH 174/829] =?UTF-8?q?feat:=20`Measure.count=20s=20=3D=200=20?= =?UTF-8?q?=E2=86=94=20s=20=3D=20=E2=88=85`=20unconditionally=20(#19287)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit By moving cases around, we can remove all side conditions from the lemma stating `count s = 0 ↔ s = ∅`. As a result, a few pairs of lemmas are unified. From LeanAPAP --- Mathlib/MeasureTheory/Measure/Count.lean | 34 ++++++------------- .../MeasureTheory/Measure/MeasureSpace.lean | 14 +++++--- Mathlib/Probability/UniformOn.lean | 16 +++++++-- 3 files changed, 34 insertions(+), 30 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/Count.lean b/Mathlib/MeasureTheory/Measure/Count.lean index b91e37e621e4a..6d5d487517d6d 100644 --- a/Mathlib/MeasureTheory/Measure/Count.lean +++ b/Mathlib/MeasureTheory/Measure/Count.lean @@ -98,33 +98,21 @@ theorem count_apply_lt_top [MeasurableSingletonClass α] : count s < ∞ ↔ s.F _ ↔ ¬s.Infinite := not_congr count_apply_eq_top _ ↔ s.Finite := Classical.not_not -theorem empty_of_count_eq_zero' (s_mble : MeasurableSet s) (hsc : count s = 0) : s = ∅ := by - have hs : s.Finite := by - rw [← count_apply_lt_top' s_mble, hsc] - exact WithTop.top_pos - simpa [count_apply_finite' hs s_mble] using hsc - -theorem empty_of_count_eq_zero [MeasurableSingletonClass α] (hsc : count s = 0) : s = ∅ := by - have hs : s.Finite := by - rw [← count_apply_lt_top, hsc] - exact WithTop.top_pos - simpa [count_apply_finite _ hs] using hsc - @[simp] -theorem count_eq_zero_iff' (s_mble : MeasurableSet s) : count s = 0 ↔ s = ∅ := - ⟨empty_of_count_eq_zero' s_mble, fun h => h.symm ▸ count_empty⟩ +theorem count_eq_zero_iff : count s = 0 ↔ s = ∅ where + mp h := eq_empty_of_forall_not_mem fun x hx ↦ by + simpa [hx] using ((ENNReal.le_tsum x).trans <| le_sum_apply _ _).trans_eq h + mpr := by rintro rfl; exact count_empty -@[simp] -theorem count_eq_zero_iff [MeasurableSingletonClass α] : count s = 0 ↔ s = ∅ := - ⟨empty_of_count_eq_zero, fun h => h.symm ▸ count_empty⟩ +lemma count_ne_zero_iff : count s ≠ 0 ↔ s.Nonempty := + count_eq_zero_iff.not.trans nonempty_iff_ne_empty.symm -theorem count_ne_zero' (hs' : s.Nonempty) (s_mble : MeasurableSet s) : count s ≠ 0 := by - rw [Ne, count_eq_zero_iff' s_mble] - exact hs'.ne_empty +alias ⟨_, count_ne_zero⟩ := count_ne_zero_iff -theorem count_ne_zero [MeasurableSingletonClass α] (hs' : s.Nonempty) : count s ≠ 0 := by - rw [Ne, count_eq_zero_iff] - exact hs'.ne_empty +@[deprecated (since := "2024-11-20")] alias ⟨empty_of_count_eq_zero, _⟩ := count_eq_zero_iff +@[deprecated (since := "2024-11-20")] alias empty_of_count_eq_zero' := empty_of_count_eq_zero +@[deprecated (since := "2024-11-20")] alias count_eq_zero_iff' := count_eq_zero_iff +@[deprecated (since := "2024-11-20")] alias count_ne_zero' := count_ne_zero @[simp] theorem count_singleton' {a : α} (ha : MeasurableSet ({a} : Set α)) : count ({a} : Set α) = 1 := by diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index 8cf14fc5dc9a6..3d62ba5101860 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -748,6 +748,11 @@ theorem coe_zero {_m : MeasurableSpace α} : ⇑(0 : Measure α) = 0 := ext s hs simp [hs] +@[simp] lemma _root_.MeasureTheory.OuterMeasure.toMeasure_eq_zero {ms : MeasurableSpace α} + {μ : OuterMeasure α} (h : ms ≤ μ.caratheodory) : μ.toMeasure h = 0 ↔ μ = 0 where + mp hμ := by ext s; exact le_bot_iff.1 <| (le_toMeasure_apply _ _ _).trans_eq congr($hμ s) + mpr := by rintro rfl; simp + @[nontriviality] lemma apply_eq_zero_of_isEmpty [IsEmpty α] {_ : MeasurableSpace α} (μ : Measure α) (s : Set α) : μ s = 0 := by @@ -1274,10 +1279,11 @@ theorem sum_apply₀ (f : ι → Measure α) {s : Set α} (hs : NullMeasurableSe /-! For the next theorem, the countability assumption is necessary. For a counterexample, consider an uncountable space, with a distinguished point `x₀`, and the sigma-algebra made of countable sets not containing `x₀`, and their complements. All points but `x₀` are measurable. -Consider the sum of the Dirac masses at points different from `x₀`, and `s = x₀`. For any Dirac mass -`δ_x`, we have `δ_x (x₀) = 0`, so `∑' x, δ_x (x₀) = 0`. On the other hand, the measure `sum δ_x` -gives mass one to each point different from `x₀`, so it gives infinite mass to any measurable set -containing `x₀` (as such a set is uncountable), and by outer regularity one get `sum δ_x {x₀} = ∞`. +Consider the sum of the Dirac masses at points different from `x₀`, and `s = {x₀}`. For any Dirac +mass `δ_x`, we have `δ_x (x₀) = 0`, so `∑' x, δ_x (x₀) = 0`. On the other hand, the measure +`sum δ_x` gives mass one to each point different from `x₀`, so it gives infinite mass to any +measurable set containing `x₀` (as such a set is uncountable), and by outer regularity one gets +`sum δ_x {x₀} = ∞`. -/ theorem sum_apply_of_countable [Countable ι] (f : ι → Measure α) (s : Set α) : sum f s = ∑' i, f i s := by diff --git a/Mathlib/Probability/UniformOn.lean b/Mathlib/Probability/UniformOn.lean index 524e33e17554c..828768e8f1610 100644 --- a/Mathlib/Probability/UniformOn.lean +++ b/Mathlib/Probability/UniformOn.lean @@ -42,7 +42,7 @@ open MeasureTheory MeasurableSpace namespace ProbabilityTheory -variable {Ω : Type*} [MeasurableSpace Ω] +variable {Ω : Type*} [MeasurableSpace Ω] {s : Set Ω} /-- Given a set `s`, `uniformOn s` is the uniform measure on `s`, defined as the counting measure conditioned by `s`. One should think of `uniformOn s t` as the proportion of `s` that is contained @@ -70,6 +70,16 @@ theorem uniformOn_empty {s : Set Ω} : uniformOn s ∅ = 0 := by simp @[deprecated (since := "2024-10-09")] alias condCount_empty := uniformOn_empty +/-- See `uniformOn_eq_zero` for a version assuming `MeasurableSingletonClass Ω` instead of +`MeasurableSet s`. -/ +@[simp] lemma uniformOn_eq_zero' (hs : MeasurableSet s) : uniformOn s = 0 ↔ s.Infinite ∨ s = ∅ := by + simp [uniformOn, hs] + +/-- See `uniformOn_eq_zero'` for a version assuming `MeasurableSet s` instead of +`MeasurableSingletonClass Ω`. -/ +@[simp] lemma uniformOn_eq_zero [MeasurableSingletonClass Ω] : + uniformOn s = 0 ↔ s.Infinite ∨ s = ∅ := by simp [uniformOn] + theorem finite_of_uniformOn_ne_zero {s t : Set Ω} (h : uniformOn s t ≠ 0) : s.Finite := by by_contra hs' simp [uniformOn, cond, Measure.count_apply_infinite hs'] at h @@ -93,7 +103,7 @@ variable [MeasurableSingletonClass Ω] theorem uniformOn_isProbabilityMeasure {s : Set Ω} (hs : s.Finite) (hs' : s.Nonempty) : IsProbabilityMeasure (uniformOn s) := by apply cond_isProbabilityMeasure_of_finite - · exact fun h => hs'.ne_empty <| Measure.empty_of_count_eq_zero h + · rwa [Measure.count_ne_zero_iff] · exact (Measure.count_apply_lt_top.2 hs).ne @[deprecated (since := "2024-10-09")] @@ -120,7 +130,7 @@ alias condCount_inter_self := uniformOn_inter_self theorem uniformOn_self (hs : s.Finite) (hs' : s.Nonempty) : uniformOn s s = 1 := by rw [uniformOn, cond_apply hs.measurableSet, Set.inter_self, ENNReal.inv_mul_cancel] - · exact fun h => hs'.ne_empty <| Measure.empty_of_count_eq_zero h + · rwa [Measure.count_ne_zero_iff] · exact (Measure.count_apply_lt_top.2 hs).ne @[deprecated (since := "2024-10-09")] From 1d199deed35dab86a152fe5ac0c3b8c8915e5369 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Iv=C3=A1n=20Renison?= <85908989+IvanRenison@users.noreply.github.com> Date: Thu, 21 Nov 2024 09:25:39 +0000 Subject: [PATCH 175/829] feat(Algebra/Ring/Parity): Add lemmas `Odd.add_one` and `Odd.one_add` (#19272) --- Mathlib/Algebra/Ring/Parity.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/Mathlib/Algebra/Ring/Parity.lean b/Mathlib/Algebra/Ring/Parity.lean index 72fdec393def4..018e1506eb3ec 100644 --- a/Mathlib/Algebra/Ring/Parity.lean +++ b/Mathlib/Algebra/Ring/Parity.lean @@ -113,6 +113,8 @@ lemma Odd.add_odd : Odd a → Odd b → Even (a + b) := by @[simp] lemma Even.add_one (h : Even a) : Odd (a + 1) := h.add_odd odd_one @[simp] lemma Even.one_add (h : Even a) : Odd (1 + a) := h.odd_add odd_one +@[simp] lemma Odd.add_one (h : Odd a) : Even (a + 1) := h.add_odd odd_one +@[simp] lemma Odd.one_add (h : Odd a) : Even (1 + a) := odd_one.add_odd h lemma odd_two_mul_add_one (a : α) : Odd (2 * a + 1) := ⟨_, rfl⟩ From e4a142c2b0da69a6c29dd6959e61e367468ead3b Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 21 Nov 2024 09:25:40 +0000 Subject: [PATCH 176/829] chore(Algebra/Group/Subgroup): split `Lattice`, `Map`, `Ker` from `Basic.lean` (#19295) The file `Algebra/Group/Subgroup/Basic.lean` is very long. Following the design of the `Submodule` folder, split it into: * `Lattice.lean`: `CompleteLattice` instance, `Subgroup.closure` * `Map.lean`: `Subgroup.map`, `Subgroup.comap` * `Ker.lean`: `MonoidHom.ker`, `MonoidHom.range` and `MonoidHom.eqLocus` --- Mathlib.lean | 3 + Mathlib/Algebra/Group/Subgroup/Basic.lean | 1410 +---------------- Mathlib/Algebra/Group/Subgroup/Ker.lean | 507 ++++++ Mathlib/Algebra/Group/Subgroup/Lattice.lean | 558 +++++++ Mathlib/Algebra/Group/Subgroup/Map.lean | 527 ++++++ .../Algebra/Group/Subgroup/ZPowers/Basic.lean | 3 +- Mathlib/Algebra/Group/Submonoid/Units.lean | 2 +- Mathlib/Algebra/Module/Submodule/Ker.lean | 1 + Mathlib/Algebra/Module/Submodule/Lattice.lean | 2 +- Mathlib/Algebra/Module/Submodule/Map.lean | 2 +- Mathlib/Analysis/Normed/Group/Basic.lean | 2 +- Mathlib/Combinatorics/Derangements/Basic.lean | 1 + Mathlib/GroupTheory/Coprod/Basic.lean | 2 +- Mathlib/GroupTheory/FreeGroup/Basic.lean | 2 +- Mathlib/GroupTheory/GroupAction/Basic.lean | 2 +- .../GroupAction/FixingSubgroup.lean | 2 +- Mathlib/GroupTheory/Perm/ClosureSwap.lean | 1 + Mathlib/GroupTheory/Perm/Sign.lean | 5 +- Mathlib/GroupTheory/PresentedGroup.lean | 1 + Mathlib/GroupTheory/QuotientGroup/Defs.lean | 2 +- Mathlib/GroupTheory/SemidirectProduct.lean | 3 +- Mathlib/GroupTheory/Subgroup/Saturated.lean | 2 +- .../RingTheory/TwoSidedIdeal/Operations.lean | 4 +- 23 files changed, 1623 insertions(+), 1421 deletions(-) create mode 100644 Mathlib/Algebra/Group/Subgroup/Ker.lean create mode 100644 Mathlib/Algebra/Group/Subgroup/Lattice.lean create mode 100644 Mathlib/Algebra/Group/Subgroup/Map.lean diff --git a/Mathlib.lean b/Mathlib.lean index 5922aa04576e9..139b1e81351a5 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -298,6 +298,9 @@ import Mathlib.Algebra.Group.Subgroup.Actions import Mathlib.Algebra.Group.Subgroup.Basic import Mathlib.Algebra.Group.Subgroup.Defs import Mathlib.Algebra.Group.Subgroup.Finite +import Mathlib.Algebra.Group.Subgroup.Ker +import Mathlib.Algebra.Group.Subgroup.Lattice +import Mathlib.Algebra.Group.Subgroup.Map import Mathlib.Algebra.Group.Subgroup.MulOpposite import Mathlib.Algebra.Group.Subgroup.MulOppositeLemmas import Mathlib.Algebra.Group.Subgroup.Order diff --git a/Mathlib/Algebra/Group/Subgroup/Basic.lean b/Mathlib/Algebra/Group/Subgroup/Basic.lean index 6fc988e05de4b..c6b11e2b701e3 100644 --- a/Mathlib/Algebra/Group/Subgroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subgroup/Basic.lean @@ -5,20 +5,13 @@ Authors: Kexing Ying -/ import Mathlib.Algebra.Group.Conj import Mathlib.Algebra.Group.Pi.Lemmas -import Mathlib.Algebra.Group.Submonoid.Operations -import Mathlib.Algebra.Group.Subgroup.Defs -import Mathlib.Data.Set.Image -import Mathlib.Tactic.ApplyFun +import Mathlib.Algebra.Group.Subgroup.Ker /-! -# Subgroups +# Basic results on subgroups -We prove subgroups of a group form a complete lattice, and results about images and preimages of -subgroups under group homomorphisms. The bundled subgroups use bundled monoid homomorphisms. - -There are also theorems about the subgroups generated by an element or a subset of a group, -defined both inductively and as the infimum of the set of subgroups containing a given -element/subset. +We prove basic results on the definitions of subgroups. The bundled subgroups use bundled monoid +homomorphisms. Special thanks goes to Amelia Livingston and Yury Kudryashov for their help and inspiration. @@ -40,31 +33,9 @@ Notation used here: Definitions in the file: -* `CompleteLattice (Subgroup G)` : the subgroups of `G` form a complete lattice - -* `Subgroup.closure k` : the minimal subgroup that includes the set `k` - -* `Subgroup.subtype` : the natural group homomorphism from a subgroup of group `G` to `G` - -* `Subgroup.gi` : `closure` forms a Galois insertion with the coercion to set - -* `Subgroup.comap H f` : the preimage of a subgroup `H` along the group homomorphism `f` is also a - subgroup - -* `Subgroup.map f H` : the image of a subgroup `H` along the group homomorphism `f` is also a - subgroup - * `Subgroup.prod H K` : the product of subgroups `H`, `K` of groups `G`, `N` respectively, `H × K` is a subgroup of `G × N` -* `MonoidHom.range f` : the range of the group homomorphism `f` is a subgroup - -* `MonoidHom.ker f` : the kernel of a group homomorphism `f` is the subgroup of elements `x : G` - such that `f x = 1` - -* `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 Subgroup inclusion is denoted `≤` rather than `⊆`, although `∈` is defined as @@ -96,43 +67,6 @@ theorem div_mem_comm_iff {a b : G} : a / b ∈ H ↔ b / a ∈ H := end SubgroupClass -/-! -### Conversion to/from `Additive`/`Multiplicative` --/ - - -section mul_add - -/-- Subgroups of a group `G` are isomorphic to additive subgroups of `Additive G`. -/ -@[simps!] -def Subgroup.toAddSubgroup : Subgroup G ≃o AddSubgroup (Additive G) where - toFun S := { Submonoid.toAddSubmonoid S.toSubmonoid with neg_mem' := S.inv_mem' } - invFun S := { AddSubmonoid.toSubmonoid S.toAddSubmonoid with inv_mem' := S.neg_mem' } - left_inv x := by cases x; rfl - right_inv x := by cases x; rfl - map_rel_iff' := Iff.rfl - -/-- Additive subgroup of an additive group `Additive G` are isomorphic to subgroup of `G`. -/ -abbrev AddSubgroup.toSubgroup' : AddSubgroup (Additive G) ≃o Subgroup G := - Subgroup.toAddSubgroup.symm - -/-- Additive subgroups of an additive group `A` are isomorphic to subgroups of `Multiplicative A`. --/ -@[simps!] -def AddSubgroup.toSubgroup : AddSubgroup A ≃o Subgroup (Multiplicative A) where - toFun S := { AddSubmonoid.toSubmonoid S.toAddSubmonoid with inv_mem' := S.neg_mem' } - invFun S := { Submonoid.toAddSubmonoid S.toSubmonoid with neg_mem' := S.inv_mem' } - left_inv x := by cases x; rfl - right_inv x := by cases x; rfl - map_rel_iff' := Iff.rfl - -/-- Subgroups of an additive group `Multiplicative A` are isomorphic to additive subgroups of `A`. --/ -abbrev Subgroup.toAddSubgroup' : Subgroup (Multiplicative A) ≃o AddSubgroup A := - AddSubgroup.toSubgroup.symm - -end mul_add - namespace Subgroup variable (H K : Subgroup G) @@ -141,691 +75,11 @@ variable (H K : Subgroup G) protected theorem div_mem_comm_iff {a b : G} : a / b ∈ H ↔ b / a ∈ H := div_mem_comm_iff -/-- The subgroup `G` of the group `G`. -/ -@[to_additive "The `AddSubgroup G` of the `AddGroup G`."] -instance : Top (Subgroup G) := - ⟨{ (⊤ : Submonoid G) with inv_mem' := fun _ => Set.mem_univ _ }⟩ - -/-- The top subgroup is isomorphic to the group. - -This is the group version of `Submonoid.topEquiv`. -/ -@[to_additive (attr := simps!) - "The top additive subgroup is isomorphic to the additive group. - - This is the additive group version of `AddSubmonoid.topEquiv`."] -def topEquiv : (⊤ : Subgroup G) ≃* G := - Submonoid.topEquiv - -/-- The trivial subgroup `{1}` of a group `G`. -/ -@[to_additive "The trivial `AddSubgroup` `{0}` of an `AddGroup` `G`."] -instance : Bot (Subgroup G) := - ⟨{ (⊥ : Submonoid G) with inv_mem' := by simp}⟩ - -@[to_additive] -instance : Inhabited (Subgroup G) := - ⟨⊥⟩ - -@[to_additive (attr := simp)] -theorem mem_bot {x : G} : x ∈ (⊥ : Subgroup G) ↔ x = 1 := - Iff.rfl - -@[to_additive (attr := simp)] -theorem mem_top (x : G) : x ∈ (⊤ : Subgroup G) := - Set.mem_univ x - -@[to_additive (attr := simp)] -theorem coe_top : ((⊤ : Subgroup G) : Set G) = Set.univ := - rfl - -@[to_additive (attr := simp)] -theorem coe_bot : ((⊥ : Subgroup G) : Set G) = {1} := - rfl - -@[to_additive] -instance : Unique (⊥ : Subgroup G) := - ⟨⟨1⟩, fun g => Subtype.ext g.2⟩ - -@[to_additive (attr := simp)] -theorem top_toSubmonoid : (⊤ : Subgroup G).toSubmonoid = ⊤ := - rfl - -@[to_additive (attr := simp)] -theorem bot_toSubmonoid : (⊥ : Subgroup G).toSubmonoid = ⊥ := - rfl - -@[to_additive] -theorem eq_bot_iff_forall : H = ⊥ ↔ ∀ x ∈ H, x = (1 : G) := - toSubmonoid_injective.eq_iff.symm.trans <| Submonoid.eq_bot_iff_forall _ - -@[to_additive] -theorem eq_bot_of_subsingleton [Subsingleton H] : H = ⊥ := by - rw [Subgroup.eq_bot_iff_forall] - intro y hy - rw [← Subgroup.coe_mk H y hy, Subsingleton.elim (⟨y, hy⟩ : H) 1, Subgroup.coe_one] - -@[to_additive (attr := simp, norm_cast)] -theorem coe_eq_univ {H : Subgroup G} : (H : Set G) = Set.univ ↔ H = ⊤ := - (SetLike.ext'_iff.trans (by rfl)).symm - -@[to_additive] -theorem coe_eq_singleton {H : Subgroup G} : (∃ g : G, (H : Set G) = {g}) ↔ H = ⊥ := - ⟨fun ⟨g, hg⟩ => - haveI : Subsingleton (H : Set G) := by - rw [hg] - infer_instance - H.eq_bot_of_subsingleton, - fun h => ⟨1, SetLike.ext'_iff.mp h⟩⟩ - -@[to_additive] -theorem nontrivial_iff_exists_ne_one (H : Subgroup G) : Nontrivial H ↔ ∃ x ∈ H, x ≠ (1 : G) := by - rw [Subtype.nontrivial_iff_exists_ne (fun x => x ∈ H) (1 : H)] - simp - -@[to_additive] -theorem exists_ne_one_of_nontrivial (H : Subgroup G) [Nontrivial H] : - ∃ x ∈ H, x ≠ 1 := by - rwa [← Subgroup.nontrivial_iff_exists_ne_one] - -@[to_additive] -theorem nontrivial_iff_ne_bot (H : Subgroup G) : Nontrivial H ↔ H ≠ ⊥ := by - rw [nontrivial_iff_exists_ne_one, ne_eq, eq_bot_iff_forall] - simp only [ne_eq, not_forall, exists_prop] - -/-- A subgroup is either the trivial subgroup or nontrivial. -/ -@[to_additive "A subgroup is either the trivial subgroup or nontrivial."] -theorem bot_or_nontrivial (H : Subgroup G) : H = ⊥ ∨ Nontrivial H := by - have := nontrivial_iff_ne_bot H - tauto - -/-- A subgroup is either the trivial subgroup or contains a non-identity element. -/ -@[to_additive "A subgroup is either the trivial subgroup or contains a nonzero element."] -theorem bot_or_exists_ne_one (H : Subgroup G) : H = ⊥ ∨ ∃ x ∈ H, x ≠ (1 : G) := by - convert H.bot_or_nontrivial - rw [nontrivial_iff_exists_ne_one] - -@[to_additive] -lemma ne_bot_iff_exists_ne_one {H : Subgroup G} : H ≠ ⊥ ↔ ∃ a : ↥H, a ≠ 1 := by - rw [← nontrivial_iff_ne_bot, nontrivial_iff_exists_ne_one] - simp only [ne_eq, Subtype.exists, mk_eq_one, exists_prop] - -/-- The inf of two subgroups is their intersection. -/ -@[to_additive "The inf of two `AddSubgroup`s is their intersection."] -instance : Min (Subgroup G) := - ⟨fun H₁ H₂ => - { H₁.toSubmonoid ⊓ H₂.toSubmonoid with - inv_mem' := fun ⟨hx, hx'⟩ => ⟨H₁.inv_mem hx, H₂.inv_mem hx'⟩ }⟩ - -@[to_additive (attr := simp)] -theorem coe_inf (p p' : Subgroup G) : ((p ⊓ p' : Subgroup G) : Set G) = (p : Set G) ∩ p' := - rfl - -@[to_additive (attr := simp)] -theorem mem_inf {p p' : Subgroup G} {x : G} : x ∈ p ⊓ p' ↔ x ∈ p ∧ x ∈ p' := - Iff.rfl - -@[to_additive] -instance : InfSet (Subgroup G) := - ⟨fun s => - { (⨅ S ∈ s, Subgroup.toSubmonoid S).copy (⋂ S ∈ s, ↑S) (by simp) with - inv_mem' := fun {x} hx => - Set.mem_biInter fun i h => i.inv_mem (by apply Set.mem_iInter₂.1 hx i h) }⟩ - -@[to_additive (attr := simp, norm_cast)] -theorem coe_sInf (H : Set (Subgroup G)) : ((sInf H : Subgroup G) : Set G) = ⋂ s ∈ H, ↑s := - rfl - -@[to_additive (attr := simp)] -theorem mem_sInf {S : Set (Subgroup G)} {x : G} : x ∈ sInf S ↔ ∀ p ∈ S, x ∈ p := - Set.mem_iInter₂ - -@[to_additive] -theorem mem_iInf {ι : Sort*} {S : ι → Subgroup G} {x : G} : (x ∈ ⨅ i, S i) ↔ ∀ i, x ∈ S i := by - simp only [iInf, mem_sInf, Set.forall_mem_range] - -@[to_additive (attr := simp, norm_cast)] -theorem coe_iInf {ι : Sort*} {S : ι → Subgroup G} : (↑(⨅ i, S i) : Set G) = ⋂ i, S i := by - simp only [iInf, coe_sInf, Set.biInter_range] - -/-- Subgroups of a group form a complete lattice. -/ -@[to_additive "The `AddSubgroup`s of an `AddGroup` form a complete lattice."] -instance : CompleteLattice (Subgroup G) := - { completeLatticeOfInf (Subgroup G) fun _s => - IsGLB.of_image SetLike.coe_subset_coe isGLB_biInf with - bot := ⊥ - bot_le := fun S _x hx => (mem_bot.1 hx).symm ▸ S.one_mem - top := ⊤ - le_top := fun _S x _hx => mem_top x - inf := (· ⊓ ·) - le_inf := fun _a _b _c ha hb _x hx => ⟨ha hx, hb hx⟩ - inf_le_left := fun _a _b _x => And.left - inf_le_right := fun _a _b _x => And.right } - -@[to_additive] -theorem mem_sup_left {S T : Subgroup G} : ∀ {x : G}, x ∈ S → x ∈ S ⊔ T := - have : S ≤ S ⊔ T := le_sup_left; fun h ↦ this h - -@[to_additive] -theorem mem_sup_right {S T : Subgroup G} : ∀ {x : G}, x ∈ T → x ∈ S ⊔ T := - have : T ≤ S ⊔ T := le_sup_right; fun h ↦ this h - -@[to_additive] -theorem mul_mem_sup {S T : Subgroup G} {x y : G} (hx : x ∈ S) (hy : y ∈ T) : x * y ∈ S ⊔ T := - (S ⊔ T).mul_mem (mem_sup_left hx) (mem_sup_right hy) - -@[to_additive] -theorem mem_iSup_of_mem {ι : Sort*} {S : ι → Subgroup G} (i : ι) : - ∀ {x : G}, x ∈ S i → x ∈ iSup S := - have : S i ≤ iSup S := le_iSup _ _; fun h ↦ this h - -@[to_additive] -theorem mem_sSup_of_mem {S : Set (Subgroup G)} {s : Subgroup G} (hs : s ∈ S) : - ∀ {x : G}, x ∈ s → x ∈ sSup S := - have : s ≤ sSup S := le_sSup hs; fun h ↦ this h - -@[to_additive (attr := simp)] -theorem subsingleton_iff : Subsingleton (Subgroup G) ↔ Subsingleton G := - ⟨fun _ => - ⟨fun x y => - have : ∀ i : G, i = 1 := fun i => - mem_bot.mp <| Subsingleton.elim (⊤ : Subgroup G) ⊥ ▸ mem_top i - (this x).trans (this y).symm⟩, - fun _ => ⟨fun x y => Subgroup.ext fun i => Subsingleton.elim 1 i ▸ by simp [Subgroup.one_mem]⟩⟩ - -@[to_additive (attr := simp)] -theorem nontrivial_iff : Nontrivial (Subgroup G) ↔ Nontrivial G := - not_iff_not.mp - ((not_nontrivial_iff_subsingleton.trans subsingleton_iff).trans - not_nontrivial_iff_subsingleton.symm) - -@[to_additive] -instance [Subsingleton G] : Unique (Subgroup G) := - ⟨⟨⊥⟩, fun a => @Subsingleton.elim _ (subsingleton_iff.mpr ‹_›) a _⟩ - -@[to_additive] -instance [Nontrivial G] : Nontrivial (Subgroup G) := - nontrivial_iff.mpr ‹_› - -@[to_additive] -theorem eq_top_iff' : H = ⊤ ↔ ∀ x : G, x ∈ H := - eq_top_iff.trans ⟨fun h m => h <| mem_top m, fun h m _ => h m⟩ - -/-- The `Subgroup` generated by a set. -/ -@[to_additive "The `AddSubgroup` generated by a set"] -def closure (k : Set G) : Subgroup G := - sInf { K | k ⊆ K } - variable {k : Set G} -@[to_additive] -theorem mem_closure {x : G} : x ∈ closure k ↔ ∀ K : Subgroup G, k ⊆ K → x ∈ K := - mem_sInf - -/-- The subgroup generated by a set includes the set. -/ -@[to_additive (attr := simp, aesop safe 20 apply (rule_sets := [SetLike])) - "The `AddSubgroup` generated by a set includes the set."] -theorem subset_closure : k ⊆ closure k := fun _ hx => mem_closure.2 fun _ hK => hK hx - -@[to_additive] -theorem not_mem_of_not_mem_closure {P : G} (hP : P ∉ closure k) : P ∉ k := fun h => - hP (subset_closure h) - open Set -/-- A subgroup `K` includes `closure k` if and only if it includes `k`. -/ -@[to_additive (attr := simp) - "An additive subgroup `K` includes `closure k` if and only if it includes `k`"] -theorem closure_le : closure k ≤ K ↔ k ⊆ K := - ⟨Subset.trans subset_closure, fun h => sInf_le h⟩ - -@[to_additive] -theorem closure_eq_of_le (h₁ : k ⊆ K) (h₂ : K ≤ closure k) : closure k = K := - le_antisymm ((closure_le <| K).2 h₁) h₂ - -/-- An induction principle for closure membership. If `p` holds for `1` and all elements of `k`, and -is preserved under multiplication and inverse, then `p` holds for all elements of the closure -of `k`. - -See also `Subgroup.closure_induction_left` and `Subgroup.closure_induction_right` for versions that -only require showing `p` is preserved by multiplication by elements in `k`. -/ -@[to_additive (attr := elab_as_elim) - "An induction principle for additive closure membership. If `p` - holds for `0` and all elements of `k`, and is preserved under addition and inverses, then `p` - holds for all elements of the additive closure of `k`. - - See also `AddSubgroup.closure_induction_left` and `AddSubgroup.closure_induction_left` for - versions that only require showing `p` is preserved by addition by elements in `k`."] -theorem closure_induction {p : (g : G) → g ∈ closure k → Prop} - (mem : ∀ x (hx : x ∈ k), p x (subset_closure hx)) (one : p 1 (one_mem _)) - (mul : ∀ x y hx hy, p x hx → p y hy → p (x * y) (mul_mem hx hy)) - (inv : ∀ x hx, p x hx → p x⁻¹ (inv_mem hx)) {x} (hx : x ∈ closure k) : p x hx := - let K : Subgroup G := - { carrier := { x | ∃ hx, p x hx } - mul_mem' := fun ⟨_, ha⟩ ⟨_, hb⟩ ↦ ⟨_, mul _ _ _ _ ha hb⟩ - one_mem' := ⟨_, one⟩ - inv_mem' := fun ⟨_, hb⟩ ↦ ⟨_, inv _ _ hb⟩ } - closure_le (K := K) |>.mpr (fun y hy ↦ ⟨subset_closure hy, mem y hy⟩) hx |>.elim fun _ ↦ id - -@[deprecated closure_induction (since := "2024-10-10")] -alias closure_induction' := closure_induction - -/-- An induction principle for closure membership for predicates with two arguments. -/ -@[to_additive (attr := elab_as_elim) - "An induction principle for additive closure membership, for - predicates with two arguments."] -theorem closure_induction₂ {p : (x y : G) → x ∈ closure k → y ∈ closure k → Prop} - (mem : ∀ (x) (y) (hx : x ∈ k) (hy : y ∈ k), p x y (subset_closure hx) (subset_closure hy)) - (one_left : ∀ x hx, p 1 x (one_mem _) hx) (one_right : ∀ x hx, p x 1 hx (one_mem _)) - (mul_left : ∀ x y z hx hy hz, p x z hx hz → p y z hy hz → p (x * y) z (mul_mem hx hy) hz) - (mul_right : ∀ y z x hy hz hx, p x y hx hy → p x z hx hz → p x (y * z) hx (mul_mem hy hz)) - (inv_left : ∀ x y hx hy, p x y hx hy → p x⁻¹ y (inv_mem hx) hy) - (inv_right : ∀ x y hx hy, p x y hx hy → p x y⁻¹ hx (inv_mem hy)) - {x y : G} (hx : x ∈ closure k) (hy : y ∈ closure k) : p x y hx hy := by - induction hy using closure_induction with - | mem z hz => induction hx using closure_induction with - | mem _ h => exact mem _ _ h hz - | one => exact one_left _ (subset_closure hz) - | mul _ _ _ _ h₁ h₂ => exact mul_left _ _ _ _ _ _ h₁ h₂ - | inv _ _ h => exact inv_left _ _ _ _ h - | one => exact one_right x hx - | mul _ _ _ _ h₁ h₂ => exact mul_right _ _ _ _ _ hx h₁ h₂ - | inv _ _ h => exact inv_right _ _ _ _ h - -@[to_additive (attr := simp)] -theorem closure_closure_coe_preimage {k : Set G} : closure (((↑) : closure k → G) ⁻¹' k) = ⊤ := - eq_top_iff.2 fun x _ ↦ Subtype.recOn x fun _ hx' ↦ - closure_induction (fun _ h ↦ subset_closure h) (one_mem _) (fun _ _ _ _ ↦ mul_mem) - (fun _ _ ↦ inv_mem) hx' - -variable (G) - -/-- `closure` forms a Galois insertion with the coercion to set. -/ -@[to_additive "`closure` forms a Galois insertion with the coercion to set."] -protected def gi : GaloisInsertion (@closure G _) (↑) where - choice s _ := closure s - gc s t := @closure_le _ _ t s - le_l_u _s := subset_closure - choice_eq _s _h := rfl - -variable {G} - -/-- Subgroup closure of a set is monotone in its argument: if `h ⊆ k`, -then `closure h ≤ closure k`. -/ -@[to_additive - "Additive subgroup closure of a set is monotone in its argument: if `h ⊆ k`, - then `closure h ≤ closure k`"] -theorem closure_mono ⦃h k : Set G⦄ (h' : h ⊆ k) : closure h ≤ closure k := - (Subgroup.gi G).gc.monotone_l h' - -/-- Closure of a subgroup `K` equals `K`. -/ -@[to_additive (attr := simp) "Additive closure of an additive subgroup `K` equals `K`"] -theorem closure_eq : closure (K : Set G) = K := - (Subgroup.gi G).l_u_eq K - -@[to_additive (attr := simp)] -theorem closure_empty : closure (∅ : Set G) = ⊥ := - (Subgroup.gi G).gc.l_bot - -@[to_additive (attr := simp)] -theorem closure_univ : closure (univ : Set G) = ⊤ := - @coe_top G _ ▸ closure_eq ⊤ - -@[to_additive] -theorem closure_union (s t : Set G) : closure (s ∪ t) = closure s ⊔ closure t := - (Subgroup.gi G).gc.l_sup - -@[to_additive] -theorem sup_eq_closure (H H' : Subgroup G) : H ⊔ H' = closure ((H : Set G) ∪ (H' : Set G)) := by - simp_rw [closure_union, closure_eq] - -@[to_additive] -theorem closure_iUnion {ι} (s : ι → Set G) : closure (⋃ i, s i) = ⨆ i, closure (s i) := - (Subgroup.gi G).gc.l_iSup - -@[to_additive (attr := simp)] -theorem closure_eq_bot_iff : closure k = ⊥ ↔ k ⊆ {1} := le_bot_iff.symm.trans <| closure_le _ - -@[to_additive] -theorem iSup_eq_closure {ι : Sort*} (p : ι → Subgroup G) : - ⨆ i, p i = closure (⋃ i, (p i : Set G)) := by simp_rw [closure_iUnion, closure_eq] - -/-- The subgroup generated by an element of a group equals the set of integer number powers of - the element. -/ -@[to_additive - "The `AddSubgroup` generated by an element of an `AddGroup` equals the set of - natural number multiples of the element."] -theorem mem_closure_singleton {x y : G} : y ∈ closure ({x} : Set G) ↔ ∃ n : ℤ, x ^ n = y := by - refine - ⟨fun hy => closure_induction ?_ ?_ ?_ ?_ hy, fun ⟨n, hn⟩ => - hn ▸ zpow_mem (subset_closure <| mem_singleton x) n⟩ - · intro y hy - rw [eq_of_mem_singleton hy] - exact ⟨1, zpow_one x⟩ - · exact ⟨0, zpow_zero x⟩ - · rintro _ _ _ _ ⟨n, rfl⟩ ⟨m, rfl⟩ - exact ⟨n + m, zpow_add x n m⟩ - rintro _ _ ⟨n, rfl⟩ - exact ⟨-n, zpow_neg x n⟩ - -@[to_additive] -theorem closure_singleton_one : closure ({1} : Set G) = ⊥ := by - simp [eq_bot_iff_forall, mem_closure_singleton] - -@[to_additive (attr := simp)] -lemma mem_closure_singleton_self (x : G) : x ∈ closure ({x} : Set G) := by - simpa [-subset_closure] using subset_closure (k := {x}) - -@[to_additive] -theorem le_closure_toSubmonoid (S : Set G) : Submonoid.closure S ≤ (closure S).toSubmonoid := - Submonoid.closure_le.2 subset_closure - -@[to_additive] -theorem closure_eq_top_of_mclosure_eq_top {S : Set G} (h : Submonoid.closure S = ⊤) : - closure S = ⊤ := - (eq_top_iff' _).2 fun _ => le_closure_toSubmonoid _ <| h.symm ▸ trivial - -@[to_additive] -theorem mem_iSup_of_directed {ι} [hι : Nonempty ι] {K : ι → Subgroup G} (hK : Directed (· ≤ ·) K) - {x : G} : x ∈ (iSup K : Subgroup G) ↔ ∃ i, x ∈ K i := by - refine ⟨?_, fun ⟨i, hi⟩ ↦ le_iSup K i hi⟩ - suffices x ∈ closure (⋃ i, (K i : Set G)) → ∃ i, x ∈ K i by - simpa only [closure_iUnion, closure_eq (K _)] using this - refine fun hx ↦ closure_induction (fun _ ↦ mem_iUnion.1) ?_ ?_ ?_ hx - · exact hι.elim fun i ↦ ⟨i, (K i).one_mem⟩ - · rintro x y _ _ ⟨i, hi⟩ ⟨j, hj⟩ - rcases hK i j with ⟨k, hki, hkj⟩ - exact ⟨k, mul_mem (hki hi) (hkj hj)⟩ - · rintro _ _ ⟨i, hi⟩ - exact ⟨i, inv_mem hi⟩ - -@[to_additive] -theorem coe_iSup_of_directed {ι} [Nonempty ι] {S : ι → Subgroup G} (hS : Directed (· ≤ ·) S) : - ((⨆ i, S i : Subgroup G) : Set G) = ⋃ i, S i := - Set.ext fun x ↦ by simp [mem_iSup_of_directed hS] - -@[to_additive] -theorem mem_sSup_of_directedOn {K : Set (Subgroup G)} (Kne : K.Nonempty) (hK : DirectedOn (· ≤ ·) K) - {x : G} : x ∈ sSup K ↔ ∃ s ∈ K, x ∈ s := by - haveI : Nonempty K := Kne.to_subtype - simp only [sSup_eq_iSup', mem_iSup_of_directed hK.directed_val, SetCoe.exists, Subtype.coe_mk, - exists_prop] - -variable {N : Type*} [Group N] {P : Type*} [Group P] - -/-- The preimage of a subgroup along a monoid homomorphism is a subgroup. -/ -@[to_additive - "The preimage of an `AddSubgroup` along an `AddMonoid` homomorphism - is an `AddSubgroup`."] -def comap {N : Type*} [Group N] (f : G →* N) (H : Subgroup N) : Subgroup G := - { H.toSubmonoid.comap f with - carrier := f ⁻¹' H - inv_mem' := fun {a} ha => show f a⁻¹ ∈ H by rw [f.map_inv]; exact H.inv_mem ha } - -@[to_additive (attr := simp)] -theorem coe_comap (K : Subgroup N) (f : G →* N) : (K.comap f : Set G) = f ⁻¹' K := - rfl - -@[simp] -theorem toAddSubgroup_comap {G₂ : Type*} [Group G₂] (f : G →* G₂) (s : Subgroup G₂) : - s.toAddSubgroup.comap (MonoidHom.toAdditive f) = Subgroup.toAddSubgroup (s.comap f) := rfl - -@[simp] -theorem _root_.AddSubgroup.toSubgroup_comap {A A₂ : Type*} [AddGroup A] [AddGroup A₂] - (f : A →+ A₂) (s : AddSubgroup A₂) : - s.toSubgroup.comap (AddMonoidHom.toMultiplicative f) = AddSubgroup.toSubgroup (s.comap f) := rfl - -@[to_additive (attr := simp)] -theorem mem_comap {K : Subgroup N} {f : G →* N} {x : G} : x ∈ K.comap f ↔ f x ∈ K := - Iff.rfl - -@[to_additive] -theorem comap_mono {f : G →* N} {K K' : Subgroup N} : K ≤ K' → comap f K ≤ comap f K' := - preimage_mono - -@[to_additive] -theorem comap_comap (K : Subgroup P) (g : N →* P) (f : G →* N) : - (K.comap g).comap f = K.comap (g.comp f) := - rfl - -@[to_additive (attr := simp)] -theorem comap_id (K : Subgroup N) : K.comap (MonoidHom.id _) = K := by - ext - rfl - -/-- The image of a subgroup along a monoid homomorphism is a subgroup. -/ -@[to_additive - "The image of an `AddSubgroup` along an `AddMonoid` homomorphism - is an `AddSubgroup`."] -def map (f : G →* N) (H : Subgroup G) : Subgroup N := - { H.toSubmonoid.map f with - carrier := f '' H - inv_mem' := by - rintro _ ⟨x, hx, rfl⟩ - exact ⟨x⁻¹, H.inv_mem hx, f.map_inv x⟩ } - -@[to_additive (attr := simp)] -theorem coe_map (f : G →* N) (K : Subgroup G) : (K.map f : Set N) = f '' K := - rfl - -@[to_additive (attr := simp)] -theorem mem_map {f : G →* N} {K : Subgroup G} {y : N} : y ∈ K.map f ↔ ∃ x ∈ K, f x = y := Iff.rfl - -@[to_additive] -theorem mem_map_of_mem (f : G →* N) {K : Subgroup G} {x : G} (hx : x ∈ K) : f x ∈ K.map f := - mem_image_of_mem f hx - -@[to_additive] -theorem apply_coe_mem_map (f : G →* N) (K : Subgroup G) (x : K) : f x ∈ K.map f := - mem_map_of_mem f x.prop - -@[to_additive] -theorem map_mono {f : G →* N} {K K' : Subgroup G} : K ≤ K' → map f K ≤ map f K' := - image_subset _ - -@[to_additive (attr := simp)] -theorem map_id : K.map (MonoidHom.id G) = K := - SetLike.coe_injective <| image_id _ - -@[to_additive] -theorem map_map (g : N →* P) (f : G →* N) : (K.map f).map g = K.map (g.comp f) := - SetLike.coe_injective <| image_image _ _ _ - -@[to_additive (attr := simp)] -theorem map_one_eq_bot : K.map (1 : G →* N) = ⊥ := - eq_bot_iff.mpr <| by - rintro x ⟨y, _, rfl⟩ - simp - -@[to_additive] -theorem mem_map_equiv {f : G ≃* N} {K : Subgroup G} {x : N} : - x ∈ K.map f.toMonoidHom ↔ f.symm x ∈ K := by - erw [@Set.mem_image_equiv _ _ (↑K) f.toEquiv x]; rfl - --- The simpNF linter says that the LHS can be simplified via `Subgroup.mem_map`. --- However this is a higher priority lemma. --- https://github.com/leanprover/std4/issues/207 -@[to_additive (attr := simp 1100, nolint simpNF)] -theorem mem_map_iff_mem {f : G →* N} (hf : Function.Injective f) {K : Subgroup G} {x : G} : - f x ∈ K.map f ↔ x ∈ K := - hf.mem_set_image - -@[to_additive] -theorem map_equiv_eq_comap_symm' (f : G ≃* N) (K : Subgroup G) : - K.map f.toMonoidHom = K.comap f.symm.toMonoidHom := - SetLike.coe_injective (f.toEquiv.image_eq_preimage K) - -@[to_additive] -theorem map_equiv_eq_comap_symm (f : G ≃* N) (K : Subgroup G) : - K.map f = K.comap (G := N) f.symm := - map_equiv_eq_comap_symm' _ _ - -@[to_additive] -theorem comap_equiv_eq_map_symm (f : N ≃* G) (K : Subgroup G) : - K.comap (G := N) f = K.map f.symm := - (map_equiv_eq_comap_symm f.symm K).symm - -@[to_additive] -theorem comap_equiv_eq_map_symm' (f : N ≃* G) (K : Subgroup G) : - K.comap f.toMonoidHom = K.map f.symm.toMonoidHom := - (map_equiv_eq_comap_symm f.symm K).symm - -@[to_additive] -theorem map_symm_eq_iff_map_eq {H : Subgroup N} {e : G ≃* N} : - H.map ↑e.symm = K ↔ K.map ↑e = H := by - constructor <;> rintro rfl - · rw [map_map, ← MulEquiv.coe_monoidHom_trans, MulEquiv.symm_trans_self, - MulEquiv.coe_monoidHom_refl, map_id] - · rw [map_map, ← MulEquiv.coe_monoidHom_trans, MulEquiv.self_trans_symm, - MulEquiv.coe_monoidHom_refl, map_id] - -@[to_additive] -theorem map_le_iff_le_comap {f : G →* N} {K : Subgroup G} {H : Subgroup N} : - K.map f ≤ H ↔ K ≤ H.comap f := - image_subset_iff - -@[to_additive] -theorem gc_map_comap (f : G →* N) : GaloisConnection (map f) (comap f) := fun _ _ => - map_le_iff_le_comap - -@[to_additive] -theorem map_sup (H K : Subgroup G) (f : G →* N) : (H ⊔ K).map f = H.map f ⊔ K.map f := - (gc_map_comap f).l_sup - -@[to_additive] -theorem map_iSup {ι : Sort*} (f : G →* N) (s : ι → Subgroup G) : - (iSup s).map f = ⨆ i, (s i).map f := - (gc_map_comap f).l_iSup - -@[to_additive] -theorem map_inf (H K : Subgroup G) (f : G →* N) (hf : Function.Injective f) : - (H ⊓ K).map f = H.map f ⊓ K.map f := SetLike.coe_injective (Set.image_inter hf) - -@[to_additive] -theorem map_iInf {ι : Sort*} [Nonempty ι] (f : G →* N) (hf : Function.Injective f) - (s : ι → Subgroup G) : (iInf s).map f = ⨅ i, (s i).map f := by - apply SetLike.coe_injective - simpa using (Set.injOn_of_injective hf).image_iInter_eq (s := SetLike.coe ∘ s) - -@[to_additive] -theorem comap_sup_comap_le (H K : Subgroup N) (f : G →* N) : - comap f H ⊔ comap f K ≤ comap f (H ⊔ K) := - Monotone.le_map_sup (fun _ _ => comap_mono) H K - -@[to_additive] -theorem iSup_comap_le {ι : Sort*} (f : G →* N) (s : ι → Subgroup N) : - ⨆ i, (s i).comap f ≤ (iSup s).comap f := - Monotone.le_map_iSup fun _ _ => comap_mono - -@[to_additive] -theorem comap_inf (H K : Subgroup N) (f : G →* N) : (H ⊓ K).comap f = H.comap f ⊓ K.comap f := - (gc_map_comap f).u_inf - -@[to_additive] -theorem comap_iInf {ι : Sort*} (f : G →* N) (s : ι → Subgroup N) : - (iInf s).comap f = ⨅ i, (s i).comap f := - (gc_map_comap f).u_iInf - -@[to_additive] -theorem map_inf_le (H K : Subgroup G) (f : G →* N) : map f (H ⊓ K) ≤ map f H ⊓ map f K := - le_inf (map_mono inf_le_left) (map_mono inf_le_right) - -@[to_additive] -theorem map_inf_eq (H K : Subgroup G) (f : G →* N) (hf : Function.Injective f) : - map f (H ⊓ K) = map f H ⊓ map f K := by - rw [← SetLike.coe_set_eq] - simp [Set.image_inter hf] - -@[to_additive (attr := simp)] -theorem map_bot (f : G →* N) : (⊥ : Subgroup G).map f = ⊥ := - (gc_map_comap f).l_bot - -@[to_additive (attr := simp)] -theorem map_top_of_surjective (f : G →* N) (h : Function.Surjective f) : Subgroup.map f ⊤ = ⊤ := by - rw [eq_top_iff] - intro x _ - obtain ⟨y, hy⟩ := h x - exact ⟨y, trivial, hy⟩ - -@[to_additive (attr := simp)] -theorem comap_top (f : G →* N) : (⊤ : Subgroup N).comap f = ⊤ := - (gc_map_comap f).u_top - -/-- For any subgroups `H` and `K`, view `H ⊓ K` as a subgroup of `K`. -/ -@[to_additive "For any subgroups `H` and `K`, view `H ⊓ K` as a subgroup of `K`."] -def subgroupOf (H K : Subgroup G) : Subgroup K := - H.comap K.subtype - -/-- If `H ≤ K`, then `H` as a subgroup of `K` is isomorphic to `H`. -/ -@[to_additive (attr := simps) "If `H ≤ K`, then `H` as a subgroup of `K` is isomorphic to `H`."] -def subgroupOfEquivOfLe {G : Type*} [Group G] {H K : Subgroup G} (h : H ≤ K) : - H.subgroupOf K ≃* H where - toFun g := ⟨g.1, g.2⟩ - invFun g := ⟨⟨g.1, h g.2⟩, g.2⟩ - left_inv _g := Subtype.ext (Subtype.ext rfl) - right_inv _g := Subtype.ext rfl - map_mul' _g _h := rfl - -@[to_additive (attr := simp)] -theorem comap_subtype (H K : Subgroup G) : H.comap K.subtype = H.subgroupOf K := - rfl - -@[to_additive (attr := simp)] -theorem comap_inclusion_subgroupOf {K₁ K₂ : Subgroup G} (h : K₁ ≤ K₂) (H : Subgroup G) : - (H.subgroupOf K₂).comap (inclusion h) = H.subgroupOf K₁ := - rfl - -@[to_additive] -theorem coe_subgroupOf (H K : Subgroup G) : (H.subgroupOf K : Set K) = K.subtype ⁻¹' H := - rfl - -@[to_additive] -theorem mem_subgroupOf {H K : Subgroup G} {h : K} : h ∈ H.subgroupOf K ↔ (h : G) ∈ H := - Iff.rfl - --- TODO(kmill): use `K ⊓ H` order for RHS to match `Subtype.image_preimage_coe` -@[to_additive (attr := simp)] -theorem subgroupOf_map_subtype (H K : Subgroup G) : (H.subgroupOf K).map K.subtype = H ⊓ K := - SetLike.ext' <| by refine Subtype.image_preimage_coe _ _ |>.trans ?_; apply Set.inter_comm - -@[to_additive (attr := simp)] -theorem bot_subgroupOf : (⊥ : Subgroup G).subgroupOf H = ⊥ := - Eq.symm (Subgroup.ext fun _g => Subtype.ext_iff) - -@[to_additive (attr := simp)] -theorem top_subgroupOf : (⊤ : Subgroup G).subgroupOf H = ⊤ := - rfl - -@[to_additive] -theorem subgroupOf_bot_eq_bot : H.subgroupOf ⊥ = ⊥ := - Subsingleton.elim _ _ - -@[to_additive] -theorem subgroupOf_bot_eq_top : H.subgroupOf ⊥ = ⊤ := - Subsingleton.elim _ _ - -@[to_additive (attr := simp)] -theorem subgroupOf_self : H.subgroupOf H = ⊤ := - top_unique fun g _hg => g.2 - -@[to_additive (attr := simp)] -theorem subgroupOf_inj {H₁ H₂ K : Subgroup G} : - H₁.subgroupOf K = H₂.subgroupOf K ↔ H₁ ⊓ K = H₂ ⊓ K := by - simpa only [SetLike.ext_iff, mem_inf, mem_subgroupOf, and_congr_left_iff] using Subtype.forall - -@[to_additive (attr := simp)] -theorem inf_subgroupOf_right (H K : Subgroup G) : (H ⊓ K).subgroupOf K = H.subgroupOf K := - subgroupOf_inj.2 (inf_right_idem _ _) - -@[to_additive (attr := simp)] -theorem inf_subgroupOf_left (H K : Subgroup G) : (K ⊓ H).subgroupOf K = H.subgroupOf K := by - rw [inf_comm, inf_subgroupOf_right] - -@[to_additive (attr := simp)] -theorem subgroupOf_eq_bot {H K : Subgroup G} : H.subgroupOf K = ⊥ ↔ Disjoint H K := by - rw [disjoint_iff, ← bot_subgroupOf, subgroupOf_inj, bot_inf_eq] - -@[to_additive (attr := simp)] -theorem subgroupOf_eq_top {H K : Subgroup G} : H.subgroupOf K = ⊤ ↔ K ≤ H := by - rw [← top_subgroupOf, subgroupOf_inj, top_inf_eq, inf_eq_right] +variable {N : Type*} [Group N] {P : Type*} [Group P] /-- Given `Subgroup`s `H`, `K` of groups `G`, `N` respectively, `H × K` as a subgroup of `G × N`. -/ @[to_additive prod @@ -1143,66 +397,8 @@ variable (H) end Normalizer -@[to_additive] -instance map_isCommutative (f : G →* G') [H.IsCommutative] : (H.map f).IsCommutative := - ⟨⟨by - rintro ⟨-, a, ha, rfl⟩ ⟨-, b, hb, rfl⟩ - rw [Subtype.ext_iff, coe_mul, coe_mul, Subtype.coe_mk, Subtype.coe_mk, ← map_mul, ← map_mul] - exact congr_arg f (Subtype.ext_iff.mp (mul_comm (⟨a, ha⟩ : H) ⟨b, hb⟩))⟩⟩ - -@[to_additive] -theorem comap_injective_isCommutative {f : G' →* G} (hf : Injective f) [H.IsCommutative] : - (H.comap f).IsCommutative := - ⟨⟨fun a b => - Subtype.ext - (by - have := mul_comm (⟨f a, a.2⟩ : H) (⟨f b, b.2⟩ : H) - rwa [Subtype.ext_iff, coe_mul, coe_mul, coe_mk, coe_mk, ← map_mul, ← map_mul, - hf.eq_iff] at this)⟩⟩ - -@[to_additive] -instance subgroupOf_isCommutative [H.IsCommutative] : (H.subgroupOf K).IsCommutative := - H.comap_injective_isCommutative Subtype.coe_injective - end Subgroup -namespace MulEquiv -variable {H : Type*} [Group H] - -/-- -An isomorphism of groups gives an order isomorphism between the lattices of subgroups, -defined by sending subgroups to their inverse images. - -See also `MulEquiv.mapSubgroup` which maps subgroups to their forward images. --/ -@[simps] -def comapSubgroup (f : G ≃* H) : Subgroup H ≃o Subgroup G where - toFun := Subgroup.comap f - invFun := Subgroup.comap f.symm - left_inv sg := by simp [Subgroup.comap_comap] - right_inv sh := by simp [Subgroup.comap_comap] - map_rel_iff' {sg1 sg2} := - ⟨fun h => by simpa [Subgroup.comap_comap] using - Subgroup.comap_mono (f := (f.symm : H →* G)) h, Subgroup.comap_mono⟩ - -/-- -An isomorphism of groups gives an order isomorphism between the lattices of subgroups, -defined by sending subgroups to their forward images. - -See also `MulEquiv.comapSubgroup` which maps subgroups to their inverse images. --/ -@[simps] -def mapSubgroup {H : Type*} [Group H] (f : G ≃* H) : Subgroup G ≃o Subgroup H where - toFun := Subgroup.map f - invFun := Subgroup.map f.symm - left_inv sg := by simp [Subgroup.map_map] - right_inv sh := by simp [Subgroup.map_map] - map_rel_iff' {sg1 sg2} := - ⟨fun h => by simpa [Subgroup.map_map] using - Subgroup.map_mono (f := (f.symm : H →* G)) h, Subgroup.map_mono⟩ - -end MulEquiv - namespace Group variable {s : Set G} @@ -1352,243 +548,10 @@ variable {N : Type*} {P : Type*} [Group N] [Group P] (K : Subgroup G) open Subgroup -/-- The range of a monoid homomorphism from a group is a subgroup. -/ -@[to_additive "The range of an `AddMonoidHom` from an `AddGroup` is an `AddSubgroup`."] -def range (f : G →* N) : Subgroup N := - Subgroup.copy ((⊤ : Subgroup G).map f) (Set.range f) (by simp [Set.ext_iff]) - -@[to_additive (attr := simp)] -theorem coe_range (f : G →* N) : (f.range : Set N) = Set.range f := - rfl - -@[to_additive (attr := simp)] -theorem mem_range {f : G →* N} {y : N} : y ∈ f.range ↔ ∃ x, f x = y := - Iff.rfl - -@[to_additive] -theorem range_eq_map (f : G →* N) : f.range = (⊤ : Subgroup G).map f := by ext; simp - -@[to_additive] -instance range_isCommutative {G : Type*} [CommGroup G] {N : Type*} [Group N] (f : G →* N) : - f.range.IsCommutative := - range_eq_map f ▸ Subgroup.map_isCommutative ⊤ f - -@[to_additive (attr := simp)] -theorem restrict_range (f : G →* N) : (f.restrict K).range = K.map f := by - simp_rw [SetLike.ext_iff, mem_range, mem_map, restrict_apply, SetLike.exists, - exists_prop, forall_const] - -/-- The canonical surjective group homomorphism `G →* f(G)` induced by a group -homomorphism `G →* N`. -/ -@[to_additive - "The canonical surjective `AddGroup` homomorphism `G →+ f(G)` induced by a group - homomorphism `G →+ N`."] -def rangeRestrict (f : G →* N) : G →* f.range := - codRestrict f _ fun x => ⟨x, rfl⟩ - -@[to_additive (attr := simp)] -theorem coe_rangeRestrict (f : G →* N) (g : G) : (f.rangeRestrict g : N) = f g := - rfl - -@[to_additive] -theorem coe_comp_rangeRestrict (f : G →* N) : - ((↑) : f.range → N) ∘ (⇑f.rangeRestrict : G → f.range) = f := - rfl - -@[to_additive] -theorem subtype_comp_rangeRestrict (f : G →* N) : f.range.subtype.comp f.rangeRestrict = f := - ext <| f.coe_rangeRestrict - -@[to_additive] -theorem rangeRestrict_surjective (f : G →* N) : Function.Surjective f.rangeRestrict := - fun ⟨_, g, rfl⟩ => ⟨g, rfl⟩ - -@[to_additive (attr := simp)] -lemma rangeRestrict_injective_iff {f : G →* N} : Injective f.rangeRestrict ↔ Injective f := by - convert Set.injective_codRestrict _ - -@[to_additive] -theorem map_range (g : N →* P) (f : G →* N) : f.range.map g = (g.comp f).range := by - rw [range_eq_map, range_eq_map]; exact (⊤ : Subgroup G).map_map g f - -@[to_additive] -theorem range_eq_top {N} [Group N] {f : G →* N} : - f.range = (⊤ : Subgroup N) ↔ Function.Surjective f := - SetLike.ext'_iff.trans <| Iff.trans (by rw [coe_range, coe_top]) Set.range_eq_univ - -@[deprecated (since := "2024-11-11")] alias range_top_iff_surjective := range_eq_top - -/-- The range of a surjective monoid homomorphism is the whole of the codomain. -/ -@[to_additive (attr := simp) - "The range of a surjective `AddMonoid` homomorphism is the whole of the codomain."] -theorem range_eq_top_of_surjective {N} [Group N] (f : G →* N) (hf : Function.Surjective f) : - f.range = (⊤ : Subgroup N) := - range_eq_top.2 hf - -@[deprecated (since := "2024-11-11")] alias range_top_of_surjective := range_eq_top_of_surjective - -@[to_additive (attr := simp)] -theorem range_one : (1 : G →* N).range = ⊥ := - SetLike.ext fun x => by simpa using @comm _ (· = ·) _ 1 x - -@[to_additive (attr := simp)] -theorem _root_.Subgroup.subtype_range (H : Subgroup G) : H.subtype.range = H := by - rw [range_eq_map, ← SetLike.coe_set_eq, coe_map, Subgroup.coeSubtype] - ext - simp - -@[to_additive (attr := simp)] -theorem _root_.Subgroup.inclusion_range {H K : Subgroup G} (h_le : H ≤ K) : - (inclusion h_le).range = H.subgroupOf K := - Subgroup.ext fun g => Set.ext_iff.mp (Set.range_inclusion h_le) g - -@[to_additive] -theorem subgroupOf_range_eq_of_le {G₁ G₂ : Type*} [Group G₁] [Group G₂] {K : Subgroup G₂} - (f : G₁ →* G₂) (h : f.range ≤ K) : - f.range.subgroupOf K = (f.codRestrict K fun x => h ⟨x, rfl⟩).range := by - ext k - refine exists_congr ?_ - simp [Subtype.ext_iff] - -@[simp] -theorem coe_toAdditive_range (f : G →* G') : - (MonoidHom.toAdditive f).range = Subgroup.toAddSubgroup f.range := rfl - -@[simp] -theorem coe_toMultiplicative_range {A A' : Type*} [AddGroup A] [AddGroup A'] (f : A →+ A') : - (AddMonoidHom.toMultiplicative f).range = AddSubgroup.toSubgroup f.range := rfl - -/-- Computable alternative to `MonoidHom.ofInjective`. -/ -@[to_additive "Computable alternative to `AddMonoidHom.ofInjective`."] -def ofLeftInverse {f : G →* N} {g : N →* G} (h : Function.LeftInverse g f) : G ≃* f.range := - { f.rangeRestrict with - toFun := f.rangeRestrict - invFun := g ∘ f.range.subtype - left_inv := h - right_inv := by - rintro ⟨x, y, rfl⟩ - apply Subtype.ext - rw [coe_rangeRestrict, Function.comp_apply, Subgroup.coeSubtype, Subtype.coe_mk, h] } - -@[to_additive (attr := simp)] -theorem ofLeftInverse_apply {f : G →* N} {g : N →* G} (h : Function.LeftInverse g f) (x : G) : - ↑(ofLeftInverse h x) = f x := - rfl - -@[to_additive (attr := simp)] -theorem ofLeftInverse_symm_apply {f : G →* N} {g : N →* G} (h : Function.LeftInverse g f) - (x : f.range) : (ofLeftInverse h).symm x = g x := - rfl - -/-- The range of an injective group homomorphism is isomorphic to its domain. -/ -@[to_additive "The range of an injective additive group homomorphism is isomorphic to its -domain."] -noncomputable def ofInjective {f : G →* N} (hf : Function.Injective f) : G ≃* f.range := - MulEquiv.ofBijective (f.codRestrict f.range fun x => ⟨x, rfl⟩) - ⟨fun _ _ h => hf (Subtype.ext_iff.mp h), by - rintro ⟨x, y, rfl⟩ - exact ⟨y, rfl⟩⟩ - -@[to_additive] -theorem ofInjective_apply {f : G →* N} (hf : Function.Injective f) {x : G} : - ↑(ofInjective hf x) = f x := - rfl - -@[to_additive (attr := simp)] -theorem apply_ofInjective_symm {f : G →* N} (hf : Function.Injective f) (x : f.range) : - f ((ofInjective hf).symm x) = x := - Subtype.ext_iff.1 <| (ofInjective hf).apply_symm_apply x - section Ker variable {M : Type*} [MulOneClass M] -/-- The multiplicative kernel of a monoid homomorphism is the subgroup of elements `x : G` such that -`f x = 1` -/ -@[to_additive - "The additive kernel of an `AddMonoid` homomorphism is the `AddSubgroup` of elements - such that `f x = 0`"] -def ker (f : G →* M) : Subgroup G := - { MonoidHom.mker f with - inv_mem' := fun {x} (hx : f x = 1) => - calc - f x⁻¹ = f x * f x⁻¹ := by rw [hx, one_mul] - _ = 1 := by rw [← map_mul, mul_inv_cancel, map_one] } - -@[to_additive (attr := simp)] -theorem mem_ker {f : G →* M} {x : G} : x ∈ f.ker ↔ f x = 1 := - Iff.rfl - -@[to_additive] -theorem div_mem_ker_iff (f : G →* N) {x y : G} : x / y ∈ ker f ↔ f x = f y := by - rw [mem_ker, map_div, div_eq_one] - -@[to_additive] -theorem coe_ker (f : G →* M) : (f.ker : Set G) = (f : G → M) ⁻¹' {1} := - rfl - -@[to_additive (attr := simp)] -theorem ker_toHomUnits {M} [Monoid M] (f : G →* M) : f.toHomUnits.ker = f.ker := by - ext x - simp [mem_ker, Units.ext_iff] - -@[to_additive] -theorem eq_iff (f : G →* M) {x y : G} : f x = f y ↔ y⁻¹ * x ∈ f.ker := by - constructor <;> intro h - · rw [mem_ker, map_mul, h, ← map_mul, inv_mul_cancel, map_one] - · rw [← one_mul x, ← mul_inv_cancel y, mul_assoc, map_mul, mem_ker.1 h, mul_one] - -@[to_additive] -instance decidableMemKer [DecidableEq M] (f : G →* M) : DecidablePred (· ∈ f.ker) := fun x => - decidable_of_iff (f x = 1) f.mem_ker - -@[to_additive] -theorem comap_ker (g : N →* P) (f : G →* N) : g.ker.comap f = (g.comp f).ker := - rfl - -@[to_additive (attr := simp)] -theorem comap_bot (f : G →* N) : (⊥ : Subgroup N).comap f = f.ker := - rfl - -@[to_additive (attr := simp)] -theorem ker_restrict (f : G →* N) : (f.restrict K).ker = f.ker.subgroupOf K := - rfl - -@[to_additive (attr := simp)] -theorem ker_codRestrict {S} [SetLike S N] [SubmonoidClass S N] (f : G →* N) (s : S) - (h : ∀ x, f x ∈ s) : (f.codRestrict s h).ker = f.ker := - SetLike.ext fun _x => Subtype.ext_iff - -@[to_additive (attr := simp)] -theorem ker_rangeRestrict (f : G →* N) : ker (rangeRestrict f) = ker f := - ker_codRestrict _ _ _ - -@[to_additive (attr := simp)] -theorem ker_one : (1 : G →* M).ker = ⊤ := - SetLike.ext fun _x => eq_self_iff_true _ - -@[to_additive (attr := simp)] -theorem ker_id : (MonoidHom.id G).ker = ⊥ := - rfl - -@[to_additive] -theorem ker_eq_bot_iff (f : G →* M) : f.ker = ⊥ ↔ Function.Injective f := - ⟨fun h x y hxy => by rwa [eq_iff, h, mem_bot, inv_mul_eq_one, eq_comm] at hxy, fun h => - bot_unique fun _ hx => h (hx.trans f.map_one.symm)⟩ - -@[to_additive (attr := simp)] -theorem _root_.Subgroup.ker_subtype (H : Subgroup G) : H.subtype.ker = ⊥ := - H.subtype.ker_eq_bot_iff.mpr Subtype.coe_injective - -@[to_additive (attr := simp)] -theorem _root_.Subgroup.ker_inclusion {H K : Subgroup G} (h : H ≤ K) : (inclusion h).ker = ⊥ := - (inclusion h).ker_eq_bot_iff.mpr (Set.inclusion_injective h) - -@[to_additive] -theorem ker_prod {M N : Type*} [MulOneClass M] [MulOneClass N] (f : G →* M) (g : G →* N) : - (f.prod g).ker = f.ker ⊓ g.ker := - SetLike.ext fun _ => Prod.mk_eq_one - @[to_additive] theorem prodMap_comap_prod {G' : Type*} {N' : Type*} [Group G'] [Group N'] (f : G →* N) (g : G' →* N') (S : Subgroup N) (S' : Subgroup N') : @@ -1600,74 +563,14 @@ theorem ker_prodMap {G' : Type*} {N' : Type*} [Group G'] [Group N'] (f : G →* (prodMap f g).ker = f.ker.prod g.ker := by rw [← comap_bot, ← comap_bot, ← comap_bot, ← prodMap_comap_prod, bot_prod_bot] -@[to_additive] -theorem range_le_ker_iff (f : G →* G') (g : G' →* G'') : f.range ≤ g.ker ↔ g.comp f = 1 := - ⟨fun h => ext fun x => h ⟨x, rfl⟩, by rintro h _ ⟨y, rfl⟩; exact DFunLike.congr_fun h y⟩ - -@[to_additive] -instance (priority := 100) normal_ker (f : G →* M) : f.ker.Normal := - ⟨fun x hx y => by - rw [mem_ker, map_mul, map_mul, mem_ker.1 hx, mul_one, map_mul_eq_one f (mul_inv_cancel y)]⟩ - @[to_additive (attr := simp)] lemma ker_fst : ker (fst G G') = .prod ⊥ ⊤ := SetLike.ext fun _ => (iff_of_eq (and_true _)).symm @[to_additive (attr := simp)] lemma ker_snd : ker (snd G G') = .prod ⊤ ⊥ := SetLike.ext fun _ => (iff_of_eq (true_and _)).symm -@[simp] -theorem coe_toAdditive_ker (f : G →* G') : - (MonoidHom.toAdditive f).ker = Subgroup.toAddSubgroup f.ker := rfl - -@[simp] -theorem coe_toMultiplicative_ker {A A' : Type*} [AddGroup A] [AddGroup A'] (f : A →+ A') : - (AddMonoidHom.toMultiplicative f).ker = AddSubgroup.toSubgroup f.ker := rfl - end Ker -section EqLocus - -variable {M : Type*} [Monoid M] - -/-- The subgroup of elements `x : G` such that `f x = g x` -/ -@[to_additive "The additive subgroup of elements `x : G` such that `f x = g x`"] -def eqLocus (f g : G →* M) : Subgroup G := - { eqLocusM f g with inv_mem' := eq_on_inv f g } - -@[to_additive (attr := simp)] -theorem eqLocus_same (f : G →* N) : f.eqLocus f = ⊤ := - SetLike.ext fun _ => eq_self_iff_true _ - -/-- If two monoid homomorphisms are equal on a set, then they are equal on its subgroup closure. -/ -@[to_additive - "If two monoid homomorphisms are equal on a set, then they are equal on its subgroup - closure."] -theorem eqOn_closure {f g : G →* M} {s : Set G} (h : Set.EqOn f g s) : Set.EqOn f g (closure s) := - show closure s ≤ f.eqLocus g from (closure_le _).2 h - -@[to_additive] -theorem eq_of_eqOn_top {f g : G →* M} (h : Set.EqOn f g (⊤ : Subgroup G)) : f = g := - ext fun _x => h trivial - -@[to_additive] -theorem eq_of_eqOn_dense {s : Set G} (hs : closure s = ⊤) {f g : G →* M} (h : s.EqOn f g) : f = g := - eq_of_eqOn_top <| hs ▸ eqOn_closure h - -end EqLocus - -@[to_additive] -theorem closure_preimage_le (f : G →* N) (s : Set N) : closure (f ⁻¹' s) ≤ (closure s).comap f := - (closure_le _).2 fun x hx => by rw [SetLike.mem_coe, mem_comap]; exact subset_closure hx - -/-- The image under a monoid homomorphism of the subgroup generated by a set equals the subgroup -generated by the image of the set. -/ -@[to_additive - "The image under an `AddMonoid` hom of the `AddSubgroup` generated by a set equals - the `AddSubgroup` generated by the image of the set."] -theorem map_closure (f : G →* N) (s : Set G) : (closure s).map f = closure (f '' s) := - Set.image_preimage.l_comm_of_u_comm (gc_map_comap f) (Subgroup.gi N).gc (Subgroup.gi G).gc - fun _ ↦ rfl - end MonoidHom namespace Subgroup @@ -1681,14 +584,6 @@ theorem Normal.map {H : Subgroup G} (h : H.Normal) (f : G →* N) (hf : Function normalizer_eq_top.2 h] exact le_normalizer_map _ -@[to_additive] -theorem map_eq_bot_iff {f : G →* N} : H.map f = ⊥ ↔ H ≤ f.ker := - (gc_map_comap f).l_eq_bot - -@[to_additive] -theorem map_eq_bot_iff_of_injective {f : G →* N} (hf : Function.Injective f) : - H.map f = ⊥ ↔ H = ⊥ := by rw [map_eq_bot_iff, f.ker_eq_bot_iff.mpr hf, le_bot_iff] - end Subgroup namespace Subgroup @@ -1697,166 +592,6 @@ open MonoidHom variable {N : Type*} [Group N] (f : G →* N) -@[to_additive] -theorem map_le_range (H : Subgroup G) : map f H ≤ f.range := - (range_eq_map f).symm ▸ map_mono le_top - -@[to_additive] -theorem map_subtype_le {H : Subgroup G} (K : Subgroup H) : K.map H.subtype ≤ H := - (K.map_le_range H.subtype).trans (le_of_eq H.subtype_range) - -@[to_additive] -theorem ker_le_comap (H : Subgroup N) : f.ker ≤ comap f H := - comap_bot f ▸ comap_mono bot_le - -@[to_additive] -theorem map_comap_le (H : Subgroup N) : map f (comap f H) ≤ H := - (gc_map_comap f).l_u_le _ - -@[to_additive] -theorem le_comap_map (H : Subgroup G) : H ≤ comap f (map f H) := - (gc_map_comap f).le_u_l _ - -@[to_additive] -theorem map_comap_eq (H : Subgroup N) : map f (comap f H) = f.range ⊓ H := - SetLike.ext' <| by - rw [coe_map, coe_comap, Set.image_preimage_eq_inter_range, coe_inf, coe_range, Set.inter_comm] - -@[to_additive] -theorem comap_map_eq (H : Subgroup G) : comap f (map f H) = H ⊔ f.ker := by - refine le_antisymm ?_ (sup_le (le_comap_map _ _) (ker_le_comap _ _)) - intro x hx; simp only [exists_prop, mem_map, mem_comap] at hx - rcases hx with ⟨y, hy, hy'⟩ - rw [← mul_inv_cancel_left y x] - exact mul_mem_sup hy (by simp [mem_ker, hy']) - -@[to_additive] -theorem map_comap_eq_self {f : G →* N} {H : Subgroup N} (h : H ≤ f.range) : - map f (comap f H) = H := by - rwa [map_comap_eq, inf_eq_right] - -@[to_additive] -theorem map_comap_eq_self_of_surjective {f : G →* N} (h : Function.Surjective f) (H : Subgroup N) : - map f (comap f H) = H := - map_comap_eq_self (range_eq_top.2 h ▸ le_top) - -@[to_additive] -theorem comap_le_comap_of_le_range {f : G →* N} {K L : Subgroup N} (hf : K ≤ f.range) : - K.comap f ≤ L.comap f ↔ K ≤ L := - ⟨(map_comap_eq_self hf).ge.trans ∘ map_le_iff_le_comap.mpr, comap_mono⟩ - -@[to_additive] -theorem comap_le_comap_of_surjective {f : G →* N} {K L : Subgroup N} (hf : Function.Surjective f) : - K.comap f ≤ L.comap f ↔ K ≤ L := - comap_le_comap_of_le_range (range_eq_top.2 hf ▸ le_top) - -@[to_additive] -theorem comap_lt_comap_of_surjective {f : G →* N} {K L : Subgroup N} (hf : Function.Surjective f) : - K.comap f < L.comap f ↔ K < L := by simp_rw [lt_iff_le_not_le, comap_le_comap_of_surjective hf] - -@[to_additive] -theorem comap_injective {f : G →* N} (h : Function.Surjective f) : Function.Injective (comap f) := - fun K L => by simp only [le_antisymm_iff, comap_le_comap_of_surjective h, imp_self] - -@[to_additive] -theorem comap_map_eq_self {f : G →* N} {H : Subgroup G} (h : f.ker ≤ H) : - comap f (map f H) = H := by - rwa [comap_map_eq, sup_eq_left] - -@[to_additive] -theorem comap_map_eq_self_of_injective {f : G →* N} (h : Function.Injective f) (H : Subgroup G) : - comap f (map f H) = H := - comap_map_eq_self (((ker_eq_bot_iff _).mpr h).symm ▸ bot_le) - -@[to_additive] -theorem map_le_map_iff {f : G →* N} {H K : Subgroup G} : H.map f ≤ K.map f ↔ H ≤ K ⊔ f.ker := by - rw [map_le_iff_le_comap, comap_map_eq] - -@[to_additive] -theorem map_le_map_iff' {f : G →* N} {H K : Subgroup G} : - H.map f ≤ K.map f ↔ H ⊔ f.ker ≤ K ⊔ f.ker := by - simp only [map_le_map_iff, sup_le_iff, le_sup_right, and_true] - -@[to_additive] -theorem map_eq_map_iff {f : G →* N} {H K : Subgroup G} : - H.map f = K.map f ↔ H ⊔ f.ker = K ⊔ f.ker := by simp only [le_antisymm_iff, map_le_map_iff'] - -@[to_additive] -theorem map_eq_range_iff {f : G →* N} {H : Subgroup G} : - H.map f = f.range ↔ Codisjoint H f.ker := by - rw [f.range_eq_map, map_eq_map_iff, codisjoint_iff, top_sup_eq] - -@[to_additive] -theorem map_le_map_iff_of_injective {f : G →* N} (hf : Function.Injective f) {H K : Subgroup G} : - H.map f ≤ K.map f ↔ H ≤ K := by rw [map_le_iff_le_comap, comap_map_eq_self_of_injective hf] - -@[to_additive (attr := simp)] -theorem map_subtype_le_map_subtype {G' : Subgroup G} {H K : Subgroup G'} : - H.map G'.subtype ≤ K.map G'.subtype ↔ H ≤ K := - map_le_map_iff_of_injective <| by apply Subtype.coe_injective - -@[to_additive] -theorem map_injective {f : G →* N} (h : Function.Injective f) : Function.Injective (map f) := - Function.LeftInverse.injective <| comap_map_eq_self_of_injective h - -@[to_additive] -theorem map_eq_comap_of_inverse {f : G →* N} {g : N →* G} (hl : Function.LeftInverse g f) - (hr : Function.RightInverse g f) (H : Subgroup G) : map f H = comap g H := - SetLike.ext' <| by rw [coe_map, coe_comap, Set.image_eq_preimage_of_inverse hl hr] - -/-- Given `f(A) = f(B)`, `ker f ≤ A`, and `ker f ≤ B`, deduce that `A = B`. -/ -@[to_additive "Given `f(A) = f(B)`, `ker f ≤ A`, and `ker f ≤ B`, deduce that `A = B`."] -theorem map_injective_of_ker_le {H K : Subgroup G} (hH : f.ker ≤ H) (hK : f.ker ≤ K) - (hf : map f H = map f K) : H = K := by - apply_fun comap f at hf - rwa [comap_map_eq, comap_map_eq, sup_of_le_left hH, sup_of_le_left hK] at hf - -@[to_additive] -theorem closure_preimage_eq_top (s : Set G) : closure ((closure s).subtype ⁻¹' s) = ⊤ := by - apply map_injective (closure s).subtype_injective - rw [MonoidHom.map_closure, ← MonoidHom.range_eq_map, subtype_range, - Set.image_preimage_eq_of_subset] - rw [coeSubtype, Subtype.range_coe_subtype] - exact subset_closure - -@[to_additive] -theorem comap_sup_eq_of_le_range {H K : Subgroup N} (hH : H ≤ f.range) (hK : K ≤ f.range) : - comap f H ⊔ comap f K = comap f (H ⊔ K) := - map_injective_of_ker_le f ((ker_le_comap f H).trans le_sup_left) (ker_le_comap f (H ⊔ K)) - (by - rw [map_comap_eq, map_sup, map_comap_eq, map_comap_eq, inf_eq_right.mpr hH, - inf_eq_right.mpr hK, inf_eq_right.mpr (sup_le hH hK)]) - -@[to_additive] -theorem comap_sup_eq (H K : Subgroup N) (hf : Function.Surjective f) : - comap f H ⊔ comap f K = comap f (H ⊔ K) := - comap_sup_eq_of_le_range f (range_eq_top.2 hf ▸ le_top) (range_eq_top.2 hf ▸ le_top) - -@[to_additive] -theorem sup_subgroupOf_eq {H K L : Subgroup G} (hH : H ≤ L) (hK : K ≤ L) : - H.subgroupOf L ⊔ K.subgroupOf L = (H ⊔ K).subgroupOf L := - comap_sup_eq_of_le_range L.subtype (hH.trans L.subtype_range.ge) (hK.trans L.subtype_range.ge) - -@[to_additive] -theorem codisjoint_subgroupOf_sup (H K : Subgroup G) : - Codisjoint (H.subgroupOf (H ⊔ K)) (K.subgroupOf (H ⊔ K)) := by - rw [codisjoint_iff, sup_subgroupOf_eq, subgroupOf_self] - exacts [le_sup_left, le_sup_right] - -/-- A subgroup is isomorphic to its image under an injective function. If you have an isomorphism, -use `MulEquiv.subgroupMap` for better definitional equalities. -/ -@[to_additive - "An additive subgroup is isomorphic to its image under an injective function. If you - have an isomorphism, use `AddEquiv.addSubgroupMap` for better definitional equalities."] -noncomputable def equivMapOfInjective (H : Subgroup G) (f : G →* N) (hf : Function.Injective f) : - H ≃* H.map f := - { Equiv.Set.image f H hf with map_mul' := fun _ _ => Subtype.ext (f.map_mul _ _) } - -@[to_additive (attr := simp)] -theorem coe_equivMapOfInjective_apply (H : Subgroup G) (f : G →* N) (hf : Function.Injective f) - (h : H) : (equivMapOfInjective H f hf h : N) = f h := - rfl - /-- The preimage of the normalizer is equal to the normalizer of the preimage of a surjective function. -/ @[to_additive @@ -2064,103 +799,6 @@ theorem Normal.of_map_subtype {K : Subgroup G} {L : Subgroup K} end Subgroup -namespace MonoidHom - -/-- The `MonoidHom` from the preimage of a subgroup to itself. -/ -@[to_additive (attr := simps!) "the `AddMonoidHom` from the preimage of an -additive subgroup to itself."] -def subgroupComap (f : G →* G') (H' : Subgroup G') : H'.comap f →* H' := - f.submonoidComap H'.toSubmonoid - -/-- The `MonoidHom` from a subgroup to its image. -/ -@[to_additive (attr := simps!) "the `AddMonoidHom` from an additive subgroup to its image"] -def subgroupMap (f : G →* G') (H : Subgroup G) : H →* H.map f := - f.submonoidMap H.toSubmonoid - -@[to_additive] -theorem subgroupMap_surjective (f : G →* G') (H : Subgroup G) : - Function.Surjective (f.subgroupMap H) := - f.submonoidMap_surjective H.toSubmonoid - -end MonoidHom - -namespace MulEquiv - -variable {H K : Subgroup G} - -/-- Makes the identity isomorphism from a proof two subgroups of a multiplicative - group are equal. -/ -@[to_additive - "Makes the identity additive isomorphism from a proof - two subgroups of an additive group are equal."] -def subgroupCongr (h : H = K) : H ≃* K := - { Equiv.setCongr <| congr_arg _ h with map_mul' := fun _ _ => rfl } - -@[to_additive (attr := simp)] -lemma subgroupCongr_apply (h : H = K) (x) : - (MulEquiv.subgroupCongr h x : G) = x := rfl - -@[to_additive (attr := simp)] -lemma subgroupCongr_symm_apply (h : H = K) (x) : - ((MulEquiv.subgroupCongr h).symm x : G) = x := rfl - -/-- A subgroup is isomorphic to its image under an isomorphism. If you only have an injective map, -use `Subgroup.equiv_map_of_injective`. -/ -@[to_additive - "An additive subgroup is isomorphic to its image under an isomorphism. If you only - have an injective map, use `AddSubgroup.equiv_map_of_injective`."] -def subgroupMap (e : G ≃* G') (H : Subgroup G) : H ≃* H.map (e : G →* G') := - MulEquiv.submonoidMap (e : G ≃* G') H.toSubmonoid - -@[to_additive (attr := simp)] -theorem coe_subgroupMap_apply (e : G ≃* G') (H : Subgroup G) (g : H) : - ((subgroupMap e H g : H.map (e : G →* G')) : G') = e g := - rfl - -@[to_additive (attr := simp)] -theorem subgroupMap_symm_apply (e : G ≃* G') (H : Subgroup G) (g : H.map (e : G →* G')) : - (e.subgroupMap H).symm g = ⟨e.symm g, SetLike.mem_coe.1 <| Set.mem_image_equiv.1 g.2⟩ := - rfl - -end MulEquiv - -namespace Subgroup - -@[to_additive (attr := simp)] -theorem equivMapOfInjective_coe_mulEquiv (H : Subgroup G) (e : G ≃* G') : - H.equivMapOfInjective (e : G →* G') (EquivLike.injective e) = e.subgroupMap H := by - ext - rfl - -variable {C : Type*} [CommGroup C] {s t : Subgroup C} {x : C} - -@[to_additive] -theorem mem_sup : x ∈ s ⊔ t ↔ ∃ y ∈ s, ∃ z ∈ t, y * z = x := - ⟨fun h => by - rw [sup_eq_closure] at h - refine Subgroup.closure_induction ?_ ?_ ?_ ?_ h - · rintro y (h | h) - · exact ⟨y, h, 1, t.one_mem, by simp⟩ - · exact ⟨1, s.one_mem, y, h, by simp⟩ - · exact ⟨1, s.one_mem, 1, ⟨t.one_mem, mul_one 1⟩⟩ - · rintro _ _ _ _ ⟨y₁, hy₁, z₁, hz₁, rfl⟩ ⟨y₂, hy₂, z₂, hz₂, rfl⟩ - exact ⟨_, mul_mem hy₁ hy₂, _, mul_mem hz₁ hz₂, by simp [mul_assoc, mul_left_comm]⟩ - · rintro _ _ ⟨y, hy, z, hz, rfl⟩ - exact ⟨_, inv_mem hy, _, inv_mem hz, mul_comm z y ▸ (mul_inv_rev z y).symm⟩, by - rintro ⟨y, hy, z, hz, rfl⟩; exact mul_mem_sup hy hz⟩ - -@[to_additive] -theorem mem_sup' : x ∈ s ⊔ t ↔ ∃ (y : s) (z : t), (y : C) * z = x := - mem_sup.trans <| by simp only [SetLike.exists, coe_mk, exists_prop] - -@[to_additive] -theorem mem_closure_pair {x y z : C} : - z ∈ closure ({x, y} : Set C) ↔ ∃ m n : ℤ, x ^ m * y ^ n = z := by - rw [← Set.singleton_union, Subgroup.closure_union, mem_sup] - simp_rw [mem_closure_singleton, exists_exists_eq_and] - -end Subgroup - namespace Subgroup section SubgroupNormal @@ -2208,15 +846,6 @@ theorem inf_subgroupOf_inf_normal_of_left {A' A : Subgroup G} (B : Subgroup G) ( instance normal_inf_normal (H K : Subgroup G) [hH : H.Normal] [hK : K.Normal] : (H ⊓ K).Normal := ⟨fun n hmem g => ⟨hH.conj_mem n hmem.1 g, hK.conj_mem n hmem.2 g⟩⟩ -@[to_additive] -theorem subgroupOf_sup (A A' B : Subgroup G) (hA : A ≤ B) (hA' : A' ≤ B) : - (A ⊔ A').subgroupOf B = A.subgroupOf B ⊔ A'.subgroupOf B := by - refine - map_injective_of_ker_le B.subtype (ker_le_comap _ _) - (le_trans (ker_le_comap B.subtype _) le_sup_left) ?_ - simp only [subgroupOf, map_comap_eq, map_sup, subtype_range] - rw [inf_of_le_right (sup_le hA hA'), inf_of_le_right hA', inf_of_le_right hA] - @[to_additive] theorem SubgroupNormal.mem_comm {H K : Subgroup G} (hK : H ≤ K) [hN : (H.subgroupOf K).Normal] {a b : G} (hb : b ∈ K) (h : a * b ∈ H) : b * a ∈ H := by @@ -2243,33 +872,6 @@ theorem commute_of_normal_of_disjoint (H₁ H₂ : Subgroup G) (hH₁ : H₁.Nor end SubgroupNormal -@[to_additive] -theorem disjoint_def {H₁ H₂ : Subgroup G} : Disjoint H₁ H₂ ↔ ∀ {x : G}, x ∈ H₁ → x ∈ H₂ → x = 1 := - disjoint_iff_inf_le.trans <| by simp only [Disjoint, SetLike.le_def, mem_inf, mem_bot, and_imp] - -@[to_additive] -theorem disjoint_def' {H₁ H₂ : Subgroup G} : - Disjoint H₁ H₂ ↔ ∀ {x y : G}, x ∈ H₁ → y ∈ H₂ → x = y → x = 1 := - disjoint_def.trans ⟨fun h _x _y hx hy hxy ↦ h hx <| hxy.symm ▸ hy, fun h _x hx hx' ↦ h hx hx' rfl⟩ - -@[to_additive] -theorem disjoint_iff_mul_eq_one {H₁ H₂ : Subgroup G} : - Disjoint H₁ H₂ ↔ ∀ {x y : G}, x ∈ H₁ → y ∈ H₂ → x * y = 1 → x = 1 ∧ y = 1 := - disjoint_def'.trans - ⟨fun h x y hx hy hxy => - let hx1 : x = 1 := h hx (H₂.inv_mem hy) (eq_inv_iff_mul_eq_one.mpr hxy) - ⟨hx1, by simpa [hx1] using hxy⟩, - fun h _ _ hx hy hxy => (h hx (H₂.inv_mem hy) (mul_inv_eq_one.mpr hxy)).1⟩ - -@[to_additive] -theorem mul_injective_of_disjoint {H₁ H₂ : Subgroup G} (h : Disjoint H₁ H₂) : - Function.Injective (fun g => g.1 * g.2 : H₁ × H₂ → G) := by - intro x y hxy - rw [← inv_mul_eq_iff_eq_mul, ← mul_assoc, ← mul_inv_eq_one, mul_assoc] at hxy - replace hxy := disjoint_iff_mul_eq_one.mp h (y.1⁻¹ * x.1).prop (x.2 * y.2⁻¹).prop hxy - rwa [coe_mul, coe_mul, coe_inv, coe_inv, inv_mul_eq_one, mul_inv_eq_one, ← Subtype.ext_iff, ← - Subtype.ext_iff, eq_comm, ← Prod.ext_iff] at hxy - end Subgroup namespace IsConj @@ -2310,5 +912,3 @@ def noncenter (G : Type*) [Monoid G] : Set (ConjClasses G) := g ∈ noncenter G ↔ g.carrier.Nontrivial := Iff.rfl end ConjClasses - -set_option linter.style.longFile 2500 diff --git a/Mathlib/Algebra/Group/Subgroup/Ker.lean b/Mathlib/Algebra/Group/Subgroup/Ker.lean new file mode 100644 index 0000000000000..8b9be58076933 --- /dev/null +++ b/Mathlib/Algebra/Group/Subgroup/Ker.lean @@ -0,0 +1,507 @@ +/- +Copyright (c) 2020 Kexing Ying. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kexing Ying +-/ +import Mathlib.Algebra.Group.Subgroup.Map +import Mathlib.Tactic.ApplyFun + +/-! +# Kernel and range of group homomorphisms + +We define and prove results about the kernel and range of group homomorphisms. + +Special thanks goes to Amelia Livingston and Yury Kudryashov for their help and inspiration. + +## Main definitions + +Notation used here: + +- `G N` are `Group`s + +- `x` is an element of type `G` + +- `f g : N →* G` are group homomorphisms + +Definitions in the file: + +* `MonoidHom.range f` : the range of the group homomorphism `f` is a subgroup + +* `MonoidHom.ker f` : the kernel of a group homomorphism `f` is the subgroup of elements `x : G` + such that `f x = 1` + +* `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 + +Subgroup inclusion is denoted `≤` rather than `⊆`, although `∈` is defined as +membership of a subgroup's underlying set. + +## Tags +subgroup, subgroups +-/ + +assert_not_exists OrderedAddCommMonoid +assert_not_exists Multiset +assert_not_exists Ring + +open Function +open scoped Int + +variable {G G' G'' : Type*} [Group G] [Group G'] [Group G''] +variable {A : Type*} [AddGroup A] + +namespace MonoidHom + +variable {N : Type*} {P : Type*} [Group N] [Group P] (K : Subgroup G) + +open Subgroup + +/-- The range of a monoid homomorphism from a group is a subgroup. -/ +@[to_additive "The range of an `AddMonoidHom` from an `AddGroup` is an `AddSubgroup`."] +def range (f : G →* N) : Subgroup N := + Subgroup.copy ((⊤ : Subgroup G).map f) (Set.range f) (by simp [Set.ext_iff]) + +@[to_additive (attr := simp)] +theorem coe_range (f : G →* N) : (f.range : Set N) = Set.range f := + rfl + +@[to_additive (attr := simp)] +theorem mem_range {f : G →* N} {y : N} : y ∈ f.range ↔ ∃ x, f x = y := + Iff.rfl + +@[to_additive] +theorem range_eq_map (f : G →* N) : f.range = (⊤ : Subgroup G).map f := by ext; simp + +@[to_additive] +instance range_isCommutative {G : Type*} [CommGroup G] {N : Type*} [Group N] (f : G →* N) : + f.range.IsCommutative := + range_eq_map f ▸ Subgroup.map_isCommutative ⊤ f + +@[to_additive (attr := simp)] +theorem restrict_range (f : G →* N) : (f.restrict K).range = K.map f := by + simp_rw [SetLike.ext_iff, mem_range, mem_map, restrict_apply, SetLike.exists, + exists_prop, forall_const] + +/-- The canonical surjective group homomorphism `G →* f(G)` induced by a group +homomorphism `G →* N`. -/ +@[to_additive + "The canonical surjective `AddGroup` homomorphism `G →+ f(G)` induced by a group + homomorphism `G →+ N`."] +def rangeRestrict (f : G →* N) : G →* f.range := + codRestrict f _ fun x => ⟨x, rfl⟩ + +@[to_additive (attr := simp)] +theorem coe_rangeRestrict (f : G →* N) (g : G) : (f.rangeRestrict g : N) = f g := + rfl + +@[to_additive] +theorem coe_comp_rangeRestrict (f : G →* N) : + ((↑) : f.range → N) ∘ (⇑f.rangeRestrict : G → f.range) = f := + rfl + +@[to_additive] +theorem subtype_comp_rangeRestrict (f : G →* N) : f.range.subtype.comp f.rangeRestrict = f := + ext <| f.coe_rangeRestrict + +@[to_additive] +theorem rangeRestrict_surjective (f : G →* N) : Function.Surjective f.rangeRestrict := + fun ⟨_, g, rfl⟩ => ⟨g, rfl⟩ + +@[to_additive (attr := simp)] +lemma rangeRestrict_injective_iff {f : G →* N} : Injective f.rangeRestrict ↔ Injective f := by + convert Set.injective_codRestrict _ + +@[to_additive] +theorem map_range (g : N →* P) (f : G →* N) : f.range.map g = (g.comp f).range := by + rw [range_eq_map, range_eq_map]; exact (⊤ : Subgroup G).map_map g f + +@[to_additive] +theorem range_eq_top {N} [Group N] {f : G →* N} : + f.range = (⊤ : Subgroup N) ↔ Function.Surjective f := + SetLike.ext'_iff.trans <| Iff.trans (by rw [coe_range, coe_top]) Set.range_eq_univ + +@[deprecated (since := "2024-11-11")] alias range_top_iff_surjective := range_eq_top + +/-- The range of a surjective monoid homomorphism is the whole of the codomain. -/ +@[to_additive (attr := simp) + "The range of a surjective `AddMonoid` homomorphism is the whole of the codomain."] +theorem range_eq_top_of_surjective {N} [Group N] (f : G →* N) (hf : Function.Surjective f) : + f.range = (⊤ : Subgroup N) := + range_eq_top.2 hf + +@[deprecated (since := "2024-11-11")] alias range_top_of_surjective := range_eq_top_of_surjective + +@[to_additive (attr := simp)] +theorem range_one : (1 : G →* N).range = ⊥ := + SetLike.ext fun x => by simpa using @comm _ (· = ·) _ 1 x + +@[to_additive (attr := simp)] +theorem _root_.Subgroup.subtype_range (H : Subgroup G) : H.subtype.range = H := by + rw [range_eq_map, ← SetLike.coe_set_eq, coe_map, Subgroup.coeSubtype] + ext + simp + +@[to_additive (attr := simp)] +theorem _root_.Subgroup.inclusion_range {H K : Subgroup G} (h_le : H ≤ K) : + (inclusion h_le).range = H.subgroupOf K := + Subgroup.ext fun g => Set.ext_iff.mp (Set.range_inclusion h_le) g + +@[to_additive] +theorem subgroupOf_range_eq_of_le {G₁ G₂ : Type*} [Group G₁] [Group G₂] {K : Subgroup G₂} + (f : G₁ →* G₂) (h : f.range ≤ K) : + f.range.subgroupOf K = (f.codRestrict K fun x => h ⟨x, rfl⟩).range := by + ext k + refine exists_congr ?_ + simp [Subtype.ext_iff] + +/-- Computable alternative to `MonoidHom.ofInjective`. -/ +@[to_additive "Computable alternative to `AddMonoidHom.ofInjective`."] +def ofLeftInverse {f : G →* N} {g : N →* G} (h : Function.LeftInverse g f) : G ≃* f.range := + { f.rangeRestrict with + toFun := f.rangeRestrict + invFun := g ∘ f.range.subtype + left_inv := h + right_inv := by + rintro ⟨x, y, rfl⟩ + apply Subtype.ext + rw [coe_rangeRestrict, Function.comp_apply, Subgroup.coeSubtype, Subtype.coe_mk, h] } + +@[to_additive (attr := simp)] +theorem ofLeftInverse_apply {f : G →* N} {g : N →* G} (h : Function.LeftInverse g f) (x : G) : + ↑(ofLeftInverse h x) = f x := + rfl + +@[to_additive (attr := simp)] +theorem ofLeftInverse_symm_apply {f : G →* N} {g : N →* G} (h : Function.LeftInverse g f) + (x : f.range) : (ofLeftInverse h).symm x = g x := + rfl + +/-- The range of an injective group homomorphism is isomorphic to its domain. -/ +@[to_additive "The range of an injective additive group homomorphism is isomorphic to its +domain."] +noncomputable def ofInjective {f : G →* N} (hf : Function.Injective f) : G ≃* f.range := + MulEquiv.ofBijective (f.codRestrict f.range fun x => ⟨x, rfl⟩) + ⟨fun _ _ h => hf (Subtype.ext_iff.mp h), by + rintro ⟨x, y, rfl⟩ + exact ⟨y, rfl⟩⟩ + +@[to_additive] +theorem ofInjective_apply {f : G →* N} (hf : Function.Injective f) {x : G} : + ↑(ofInjective hf x) = f x := + rfl + +@[to_additive (attr := simp)] +theorem apply_ofInjective_symm {f : G →* N} (hf : Function.Injective f) (x : f.range) : + f ((ofInjective hf).symm x) = x := + Subtype.ext_iff.1 <| (ofInjective hf).apply_symm_apply x + +@[simp] +theorem coe_toAdditive_range (f : G →* G') : + (MonoidHom.toAdditive f).range = Subgroup.toAddSubgroup f.range := rfl + +@[simp] +theorem coe_toMultiplicative_range {A A' : Type*} [AddGroup A] [AddGroup A'] (f : A →+ A') : + (AddMonoidHom.toMultiplicative f).range = AddSubgroup.toSubgroup f.range := rfl + +section Ker + +variable {M : Type*} [MulOneClass M] + +/-- The multiplicative kernel of a monoid homomorphism is the subgroup of elements `x : G` such that +`f x = 1` -/ +@[to_additive + "The additive kernel of an `AddMonoid` homomorphism is the `AddSubgroup` of elements + such that `f x = 0`"] +def ker (f : G →* M) : Subgroup G := + { MonoidHom.mker f with + inv_mem' := fun {x} (hx : f x = 1) => + calc + f x⁻¹ = f x * f x⁻¹ := by rw [hx, one_mul] + _ = 1 := by rw [← map_mul, mul_inv_cancel, map_one] } + +@[to_additive (attr := simp)] +theorem mem_ker {f : G →* M} {x : G} : x ∈ f.ker ↔ f x = 1 := + Iff.rfl + +@[to_additive] +theorem div_mem_ker_iff (f : G →* N) {x y : G} : x / y ∈ ker f ↔ f x = f y := by + rw [mem_ker, map_div, div_eq_one] + +@[to_additive] +theorem coe_ker (f : G →* M) : (f.ker : Set G) = (f : G → M) ⁻¹' {1} := + rfl + +@[to_additive (attr := simp)] +theorem ker_toHomUnits {M} [Monoid M] (f : G →* M) : f.toHomUnits.ker = f.ker := by + ext x + simp [mem_ker, Units.ext_iff] + +@[to_additive] +theorem eq_iff (f : G →* M) {x y : G} : f x = f y ↔ y⁻¹ * x ∈ f.ker := by + constructor <;> intro h + · rw [mem_ker, map_mul, h, ← map_mul, inv_mul_cancel, map_one] + · rw [← one_mul x, ← mul_inv_cancel y, mul_assoc, map_mul, mem_ker.1 h, mul_one] + +@[to_additive] +instance decidableMemKer [DecidableEq M] (f : G →* M) : DecidablePred (· ∈ f.ker) := fun x => + decidable_of_iff (f x = 1) f.mem_ker + +@[to_additive] +theorem comap_ker (g : N →* P) (f : G →* N) : g.ker.comap f = (g.comp f).ker := + rfl + +@[to_additive (attr := simp)] +theorem comap_bot (f : G →* N) : (⊥ : Subgroup N).comap f = f.ker := + rfl + +@[to_additive (attr := simp)] +theorem ker_restrict (f : G →* N) : (f.restrict K).ker = f.ker.subgroupOf K := + rfl + +@[to_additive (attr := simp)] +theorem ker_codRestrict {S} [SetLike S N] [SubmonoidClass S N] (f : G →* N) (s : S) + (h : ∀ x, f x ∈ s) : (f.codRestrict s h).ker = f.ker := + SetLike.ext fun _x => Subtype.ext_iff + +@[to_additive (attr := simp)] +theorem ker_rangeRestrict (f : G →* N) : ker (rangeRestrict f) = ker f := + ker_codRestrict _ _ _ + +@[to_additive (attr := simp)] +theorem ker_one : (1 : G →* M).ker = ⊤ := + SetLike.ext fun _x => eq_self_iff_true _ + +@[to_additive (attr := simp)] +theorem ker_id : (MonoidHom.id G).ker = ⊥ := + rfl + +@[to_additive] +theorem ker_eq_bot_iff (f : G →* M) : f.ker = ⊥ ↔ Function.Injective f := + ⟨fun h x y hxy => by rwa [eq_iff, h, mem_bot, inv_mul_eq_one, eq_comm] at hxy, fun h => + bot_unique fun _ hx => h (hx.trans f.map_one.symm)⟩ + +@[to_additive (attr := simp)] +theorem _root_.Subgroup.ker_subtype (H : Subgroup G) : H.subtype.ker = ⊥ := + H.subtype.ker_eq_bot_iff.mpr Subtype.coe_injective + +@[to_additive (attr := simp)] +theorem _root_.Subgroup.ker_inclusion {H K : Subgroup G} (h : H ≤ K) : (inclusion h).ker = ⊥ := + (inclusion h).ker_eq_bot_iff.mpr (Set.inclusion_injective h) + +@[to_additive] +theorem ker_prod {M N : Type*} [MulOneClass M] [MulOneClass N] (f : G →* M) (g : G →* N) : + (f.prod g).ker = f.ker ⊓ g.ker := + SetLike.ext fun _ => Prod.mk_eq_one + +@[to_additive] +theorem range_le_ker_iff (f : G →* G') (g : G' →* G'') : f.range ≤ g.ker ↔ g.comp f = 1 := + ⟨fun h => ext fun x => h ⟨x, rfl⟩, by rintro h _ ⟨y, rfl⟩; exact DFunLike.congr_fun h y⟩ + +@[to_additive] +instance (priority := 100) normal_ker (f : G →* M) : f.ker.Normal := + ⟨fun x hx y => by + rw [mem_ker, map_mul, map_mul, mem_ker.1 hx, mul_one, map_mul_eq_one f (mul_inv_cancel y)]⟩ + +@[simp] +theorem coe_toAdditive_ker (f : G →* G') : + (MonoidHom.toAdditive f).ker = Subgroup.toAddSubgroup f.ker := rfl + +@[simp] +theorem coe_toMultiplicative_ker {A A' : Type*} [AddGroup A] [AddGroup A'] (f : A →+ A') : + (AddMonoidHom.toMultiplicative f).ker = AddSubgroup.toSubgroup f.ker := rfl + +end Ker + +section EqLocus + +variable {M : Type*} [Monoid M] + +/-- The subgroup of elements `x : G` such that `f x = g x` -/ +@[to_additive "The additive subgroup of elements `x : G` such that `f x = g x`"] +def eqLocus (f g : G →* M) : Subgroup G := + { eqLocusM f g with inv_mem' := eq_on_inv f g } + +@[to_additive (attr := simp)] +theorem eqLocus_same (f : G →* N) : f.eqLocus f = ⊤ := + SetLike.ext fun _ => eq_self_iff_true _ + +/-- If two monoid homomorphisms are equal on a set, then they are equal on its subgroup closure. -/ +@[to_additive + "If two monoid homomorphisms are equal on a set, then they are equal on its subgroup + closure."] +theorem eqOn_closure {f g : G →* M} {s : Set G} (h : Set.EqOn f g s) : Set.EqOn f g (closure s) := + show closure s ≤ f.eqLocus g from (closure_le _).2 h + +@[to_additive] +theorem eq_of_eqOn_top {f g : G →* M} (h : Set.EqOn f g (⊤ : Subgroup G)) : f = g := + ext fun _x => h trivial + +@[to_additive] +theorem eq_of_eqOn_dense {s : Set G} (hs : closure s = ⊤) {f g : G →* M} (h : s.EqOn f g) : f = g := + eq_of_eqOn_top <| hs ▸ eqOn_closure h + +end EqLocus + +end MonoidHom + +namespace Subgroup + +variable {N : Type*} [Group N] (H : Subgroup G) + +@[to_additive] +theorem map_eq_bot_iff {f : G →* N} : H.map f = ⊥ ↔ H ≤ f.ker := + (gc_map_comap f).l_eq_bot + +@[to_additive] +theorem map_eq_bot_iff_of_injective {f : G →* N} (hf : Function.Injective f) : + H.map f = ⊥ ↔ H = ⊥ := by rw [map_eq_bot_iff, f.ker_eq_bot_iff.mpr hf, le_bot_iff] + +open MonoidHom + +variable (f : G →* N) + +@[to_additive] +theorem map_le_range (H : Subgroup G) : map f H ≤ f.range := + (range_eq_map f).symm ▸ map_mono le_top + +@[to_additive] +theorem map_subtype_le {H : Subgroup G} (K : Subgroup H) : K.map H.subtype ≤ H := + (K.map_le_range H.subtype).trans (le_of_eq H.subtype_range) + +@[to_additive] +theorem ker_le_comap (H : Subgroup N) : f.ker ≤ comap f H := + comap_bot f ▸ comap_mono bot_le + +@[to_additive] +theorem map_comap_eq (H : Subgroup N) : map f (comap f H) = f.range ⊓ H := + SetLike.ext' <| by + rw [coe_map, coe_comap, Set.image_preimage_eq_inter_range, coe_inf, coe_range, Set.inter_comm] + +@[to_additive] +theorem comap_map_eq (H : Subgroup G) : comap f (map f H) = H ⊔ f.ker := by + refine le_antisymm ?_ (sup_le (le_comap_map _ _) (ker_le_comap _ _)) + intro x hx; simp only [exists_prop, mem_map, mem_comap] at hx + rcases hx with ⟨y, hy, hy'⟩ + rw [← mul_inv_cancel_left y x] + exact mul_mem_sup hy (by simp [mem_ker, hy']) + +@[to_additive] +theorem map_comap_eq_self {f : G →* N} {H : Subgroup N} (h : H ≤ f.range) : + map f (comap f H) = H := by + rwa [map_comap_eq, inf_eq_right] + +@[to_additive] +theorem map_comap_eq_self_of_surjective {f : G →* N} (h : Function.Surjective f) (H : Subgroup N) : + map f (comap f H) = H := + map_comap_eq_self (range_eq_top.2 h ▸ le_top) + +@[to_additive] +theorem comap_le_comap_of_le_range {f : G →* N} {K L : Subgroup N} (hf : K ≤ f.range) : + K.comap f ≤ L.comap f ↔ K ≤ L := + ⟨(map_comap_eq_self hf).ge.trans ∘ map_le_iff_le_comap.mpr, comap_mono⟩ + +@[to_additive] +theorem comap_le_comap_of_surjective {f : G →* N} {K L : Subgroup N} (hf : Function.Surjective f) : + K.comap f ≤ L.comap f ↔ K ≤ L := + comap_le_comap_of_le_range (range_eq_top.2 hf ▸ le_top) + +@[to_additive] +theorem comap_lt_comap_of_surjective {f : G →* N} {K L : Subgroup N} (hf : Function.Surjective f) : + K.comap f < L.comap f ↔ K < L := by simp_rw [lt_iff_le_not_le, comap_le_comap_of_surjective hf] + +@[to_additive] +theorem comap_injective {f : G →* N} (h : Function.Surjective f) : Function.Injective (comap f) := + fun K L => by simp only [le_antisymm_iff, comap_le_comap_of_surjective h, imp_self] + +@[to_additive] +theorem comap_map_eq_self {f : G →* N} {H : Subgroup G} (h : f.ker ≤ H) : + comap f (map f H) = H := by + rwa [comap_map_eq, sup_eq_left] + +@[to_additive] +theorem comap_map_eq_self_of_injective {f : G →* N} (h : Function.Injective f) (H : Subgroup G) : + comap f (map f H) = H := + comap_map_eq_self (((ker_eq_bot_iff _).mpr h).symm ▸ bot_le) + +@[to_additive] +theorem map_le_map_iff {f : G →* N} {H K : Subgroup G} : H.map f ≤ K.map f ↔ H ≤ K ⊔ f.ker := by + rw [map_le_iff_le_comap, comap_map_eq] + +@[to_additive] +theorem map_le_map_iff' {f : G →* N} {H K : Subgroup G} : + H.map f ≤ K.map f ↔ H ⊔ f.ker ≤ K ⊔ f.ker := by + simp only [map_le_map_iff, sup_le_iff, le_sup_right, and_true] + +@[to_additive] +theorem map_eq_map_iff {f : G →* N} {H K : Subgroup G} : + H.map f = K.map f ↔ H ⊔ f.ker = K ⊔ f.ker := by simp only [le_antisymm_iff, map_le_map_iff'] + +@[to_additive] +theorem map_eq_range_iff {f : G →* N} {H : Subgroup G} : + H.map f = f.range ↔ Codisjoint H f.ker := by + rw [f.range_eq_map, map_eq_map_iff, codisjoint_iff, top_sup_eq] + +@[to_additive] +theorem map_le_map_iff_of_injective {f : G →* N} (hf : Function.Injective f) {H K : Subgroup G} : + H.map f ≤ K.map f ↔ H ≤ K := by rw [map_le_iff_le_comap, comap_map_eq_self_of_injective hf] + +@[to_additive (attr := simp)] +theorem map_subtype_le_map_subtype {G' : Subgroup G} {H K : Subgroup G'} : + H.map G'.subtype ≤ K.map G'.subtype ↔ H ≤ K := + map_le_map_iff_of_injective <| by apply Subtype.coe_injective + +@[to_additive] +theorem map_injective {f : G →* N} (h : Function.Injective f) : Function.Injective (map f) := + Function.LeftInverse.injective <| comap_map_eq_self_of_injective h + +/-- Given `f(A) = f(B)`, `ker f ≤ A`, and `ker f ≤ B`, deduce that `A = B`. -/ +@[to_additive "Given `f(A) = f(B)`, `ker f ≤ A`, and `ker f ≤ B`, deduce that `A = B`."] +theorem map_injective_of_ker_le {H K : Subgroup G} (hH : f.ker ≤ H) (hK : f.ker ≤ K) + (hf : map f H = map f K) : H = K := by + apply_fun comap f at hf + rwa [comap_map_eq, comap_map_eq, sup_of_le_left hH, sup_of_le_left hK] at hf + +@[to_additive] +theorem closure_preimage_eq_top (s : Set G) : closure ((closure s).subtype ⁻¹' s) = ⊤ := by + apply map_injective (closure s).subtype_injective + rw [MonoidHom.map_closure, ← MonoidHom.range_eq_map, subtype_range, + Set.image_preimage_eq_of_subset] + rw [coeSubtype, Subtype.range_coe_subtype] + exact subset_closure + +@[to_additive] +theorem comap_sup_eq_of_le_range {H K : Subgroup N} (hH : H ≤ f.range) (hK : K ≤ f.range) : + comap f H ⊔ comap f K = comap f (H ⊔ K) := + map_injective_of_ker_le f ((ker_le_comap f H).trans le_sup_left) (ker_le_comap f (H ⊔ K)) + (by + rw [map_comap_eq, map_sup, map_comap_eq, map_comap_eq, inf_eq_right.mpr hH, + inf_eq_right.mpr hK, inf_eq_right.mpr (sup_le hH hK)]) + +@[to_additive] +theorem comap_sup_eq (H K : Subgroup N) (hf : Function.Surjective f) : + comap f H ⊔ comap f K = comap f (H ⊔ K) := + comap_sup_eq_of_le_range f (range_eq_top.2 hf ▸ le_top) (range_eq_top.2 hf ▸ le_top) + +@[to_additive] +theorem sup_subgroupOf_eq {H K L : Subgroup G} (hH : H ≤ L) (hK : K ≤ L) : + H.subgroupOf L ⊔ K.subgroupOf L = (H ⊔ K).subgroupOf L := + comap_sup_eq_of_le_range L.subtype (hH.trans L.subtype_range.ge) (hK.trans L.subtype_range.ge) + +@[to_additive] +theorem codisjoint_subgroupOf_sup (H K : Subgroup G) : + Codisjoint (H.subgroupOf (H ⊔ K)) (K.subgroupOf (H ⊔ K)) := by + rw [codisjoint_iff, sup_subgroupOf_eq, subgroupOf_self] + exacts [le_sup_left, le_sup_right] + +@[to_additive] +theorem subgroupOf_sup (A A' B : Subgroup G) (hA : A ≤ B) (hA' : A' ≤ B) : + (A ⊔ A').subgroupOf B = A.subgroupOf B ⊔ A'.subgroupOf B := by + refine + map_injective_of_ker_le B.subtype (ker_le_comap _ _) + (le_trans (ker_le_comap B.subtype _) le_sup_left) ?_ + simp only [subgroupOf, map_comap_eq, map_sup, subtype_range] + rw [inf_of_le_right (sup_le hA hA'), inf_of_le_right hA', inf_of_le_right hA] + +end Subgroup diff --git a/Mathlib/Algebra/Group/Subgroup/Lattice.lean b/Mathlib/Algebra/Group/Subgroup/Lattice.lean new file mode 100644 index 0000000000000..ea029ac82db4b --- /dev/null +++ b/Mathlib/Algebra/Group/Subgroup/Lattice.lean @@ -0,0 +1,558 @@ +/- +Copyright (c) 2020 Kexing Ying. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kexing Ying +-/ +import Mathlib.Algebra.Group.Submonoid.Operations +import Mathlib.Algebra.Group.Subgroup.Defs + +/-! +# Lattice structure of subgroups + +We prove subgroups of a group form a complete lattice. + +There are also theorems about the subgroups generated by an element or a subset of a group, +defined both inductively and as the infimum of the set of subgroups containing a given +element/subset. + +Special thanks goes to Amelia Livingston and Yury Kudryashov for their help and inspiration. + +## Main definitions + +Notation used here: + +- `G` is a `Group` + +- `k` is a set of elements of type `G` + +Definitions in the file: + +* `CompleteLattice (Subgroup G)` : the subgroups of `G` form a complete lattice + +* `Subgroup.closure k` : the minimal subgroup that includes the set `k` + +* `Subgroup.gi` : `closure` forms a Galois insertion with the coercion to set + +## Implementation notes + +Subgroup inclusion is denoted `≤` rather than `⊆`, although `∈` is defined as +membership of a subgroup's underlying set. + +## Tags +subgroup, subgroups +-/ + +assert_not_exists OrderedAddCommMonoid +assert_not_exists Multiset +assert_not_exists Ring + +open Function +open scoped Int + +variable {G : Type*} [Group G] + +/-! +### Conversion to/from `Additive`/`Multiplicative` +-/ + +section mul_add + +variable {A : Type*} [AddGroup A] + +/-- Subgroups of a group `G` are isomorphic to additive subgroups of `Additive G`. -/ +@[simps!] +def Subgroup.toAddSubgroup : Subgroup G ≃o AddSubgroup (Additive G) where + toFun S := { Submonoid.toAddSubmonoid S.toSubmonoid with neg_mem' := S.inv_mem' } + invFun S := { AddSubmonoid.toSubmonoid S.toAddSubmonoid with inv_mem' := S.neg_mem' } + left_inv x := by cases x; rfl + right_inv x := by cases x; rfl + map_rel_iff' := Iff.rfl + +/-- Additive subgroup of an additive group `Additive G` are isomorphic to subgroup of `G`. -/ +abbrev AddSubgroup.toSubgroup' : AddSubgroup (Additive G) ≃o Subgroup G := + Subgroup.toAddSubgroup.symm + +/-- Additive subgroups of an additive group `A` are isomorphic to subgroups of `Multiplicative A`. +-/ +@[simps!] +def AddSubgroup.toSubgroup : AddSubgroup A ≃o Subgroup (Multiplicative A) where + toFun S := { AddSubmonoid.toSubmonoid S.toAddSubmonoid with inv_mem' := S.neg_mem' } + invFun S := { Submonoid.toAddSubmonoid S.toSubmonoid with neg_mem' := S.inv_mem' } + left_inv x := by cases x; rfl + right_inv x := by cases x; rfl + map_rel_iff' := Iff.rfl + +/-- Subgroups of an additive group `Multiplicative A` are isomorphic to additive subgroups of `A`. +-/ +abbrev Subgroup.toAddSubgroup' : Subgroup (Multiplicative A) ≃o AddSubgroup A := + AddSubgroup.toSubgroup.symm + +end mul_add + +namespace Subgroup + +variable (H K : Subgroup G) + +/-- The subgroup `G` of the group `G`. -/ +@[to_additive "The `AddSubgroup G` of the `AddGroup G`."] +instance : Top (Subgroup G) := + ⟨{ (⊤ : Submonoid G) with inv_mem' := fun _ => Set.mem_univ _ }⟩ + +/-- The top subgroup is isomorphic to the group. + +This is the group version of `Submonoid.topEquiv`. -/ +@[to_additive (attr := simps!) + "The top additive subgroup is isomorphic to the additive group. + + This is the additive group version of `AddSubmonoid.topEquiv`."] +def topEquiv : (⊤ : Subgroup G) ≃* G := + Submonoid.topEquiv + +/-- The trivial subgroup `{1}` of a group `G`. -/ +@[to_additive "The trivial `AddSubgroup` `{0}` of an `AddGroup` `G`."] +instance : Bot (Subgroup G) := + ⟨{ (⊥ : Submonoid G) with inv_mem' := by simp}⟩ + +@[to_additive] +instance : Inhabited (Subgroup G) := + ⟨⊥⟩ + +@[to_additive (attr := simp)] +theorem mem_bot {x : G} : x ∈ (⊥ : Subgroup G) ↔ x = 1 := + Iff.rfl + +@[to_additive (attr := simp)] +theorem mem_top (x : G) : x ∈ (⊤ : Subgroup G) := + Set.mem_univ x + +@[to_additive (attr := simp)] +theorem coe_top : ((⊤ : Subgroup G) : Set G) = Set.univ := + rfl + +@[to_additive (attr := simp)] +theorem coe_bot : ((⊥ : Subgroup G) : Set G) = {1} := + rfl + +@[to_additive] +instance : Unique (⊥ : Subgroup G) := + ⟨⟨1⟩, fun g => Subtype.ext g.2⟩ + +@[to_additive (attr := simp)] +theorem top_toSubmonoid : (⊤ : Subgroup G).toSubmonoid = ⊤ := + rfl + +@[to_additive (attr := simp)] +theorem bot_toSubmonoid : (⊥ : Subgroup G).toSubmonoid = ⊥ := + rfl + +@[to_additive] +theorem eq_bot_iff_forall : H = ⊥ ↔ ∀ x ∈ H, x = (1 : G) := + toSubmonoid_injective.eq_iff.symm.trans <| Submonoid.eq_bot_iff_forall _ + +@[to_additive] +theorem eq_bot_of_subsingleton [Subsingleton H] : H = ⊥ := by + rw [Subgroup.eq_bot_iff_forall] + intro y hy + rw [← Subgroup.coe_mk H y hy, Subsingleton.elim (⟨y, hy⟩ : H) 1, Subgroup.coe_one] + +@[to_additive (attr := simp, norm_cast)] +theorem coe_eq_univ {H : Subgroup G} : (H : Set G) = Set.univ ↔ H = ⊤ := + (SetLike.ext'_iff.trans (by rfl)).symm + +@[to_additive] +theorem coe_eq_singleton {H : Subgroup G} : (∃ g : G, (H : Set G) = {g}) ↔ H = ⊥ := + ⟨fun ⟨g, hg⟩ => + haveI : Subsingleton (H : Set G) := by + rw [hg] + infer_instance + H.eq_bot_of_subsingleton, + fun h => ⟨1, SetLike.ext'_iff.mp h⟩⟩ + +@[to_additive] +theorem nontrivial_iff_exists_ne_one (H : Subgroup G) : Nontrivial H ↔ ∃ x ∈ H, x ≠ (1 : G) := by + rw [Subtype.nontrivial_iff_exists_ne (fun x => x ∈ H) (1 : H)] + simp + +@[to_additive] +theorem exists_ne_one_of_nontrivial (H : Subgroup G) [Nontrivial H] : + ∃ x ∈ H, x ≠ 1 := by + rwa [← Subgroup.nontrivial_iff_exists_ne_one] + +@[to_additive] +theorem nontrivial_iff_ne_bot (H : Subgroup G) : Nontrivial H ↔ H ≠ ⊥ := by + rw [nontrivial_iff_exists_ne_one, ne_eq, eq_bot_iff_forall] + simp only [ne_eq, not_forall, exists_prop] + +/-- A subgroup is either the trivial subgroup or nontrivial. -/ +@[to_additive "A subgroup is either the trivial subgroup or nontrivial."] +theorem bot_or_nontrivial (H : Subgroup G) : H = ⊥ ∨ Nontrivial H := by + have := nontrivial_iff_ne_bot H + tauto + +/-- A subgroup is either the trivial subgroup or contains a non-identity element. -/ +@[to_additive "A subgroup is either the trivial subgroup or contains a nonzero element."] +theorem bot_or_exists_ne_one (H : Subgroup G) : H = ⊥ ∨ ∃ x ∈ H, x ≠ (1 : G) := by + convert H.bot_or_nontrivial + rw [nontrivial_iff_exists_ne_one] + +@[to_additive] +lemma ne_bot_iff_exists_ne_one {H : Subgroup G} : H ≠ ⊥ ↔ ∃ a : ↥H, a ≠ 1 := by + rw [← nontrivial_iff_ne_bot, nontrivial_iff_exists_ne_one] + simp only [ne_eq, Subtype.exists, mk_eq_one, exists_prop] + +/-- The inf of two subgroups is their intersection. -/ +@[to_additive "The inf of two `AddSubgroup`s is their intersection."] +instance : Min (Subgroup G) := + ⟨fun H₁ H₂ => + { H₁.toSubmonoid ⊓ H₂.toSubmonoid with + inv_mem' := fun ⟨hx, hx'⟩ => ⟨H₁.inv_mem hx, H₂.inv_mem hx'⟩ }⟩ + +@[to_additive (attr := simp)] +theorem coe_inf (p p' : Subgroup G) : ((p ⊓ p' : Subgroup G) : Set G) = (p : Set G) ∩ p' := + rfl + +@[to_additive (attr := simp)] +theorem mem_inf {p p' : Subgroup G} {x : G} : x ∈ p ⊓ p' ↔ x ∈ p ∧ x ∈ p' := + Iff.rfl + +@[to_additive] +instance : InfSet (Subgroup G) := + ⟨fun s => + { (⨅ S ∈ s, Subgroup.toSubmonoid S).copy (⋂ S ∈ s, ↑S) (by simp) with + inv_mem' := fun {x} hx => + Set.mem_biInter fun i h => i.inv_mem (by apply Set.mem_iInter₂.1 hx i h) }⟩ + +@[to_additive (attr := simp, norm_cast)] +theorem coe_sInf (H : Set (Subgroup G)) : ((sInf H : Subgroup G) : Set G) = ⋂ s ∈ H, ↑s := + rfl + +@[to_additive (attr := simp)] +theorem mem_sInf {S : Set (Subgroup G)} {x : G} : x ∈ sInf S ↔ ∀ p ∈ S, x ∈ p := + Set.mem_iInter₂ + +@[to_additive] +theorem mem_iInf {ι : Sort*} {S : ι → Subgroup G} {x : G} : (x ∈ ⨅ i, S i) ↔ ∀ i, x ∈ S i := by + simp only [iInf, mem_sInf, Set.forall_mem_range] + +@[to_additive (attr := simp, norm_cast)] +theorem coe_iInf {ι : Sort*} {S : ι → Subgroup G} : (↑(⨅ i, S i) : Set G) = ⋂ i, S i := by + simp only [iInf, coe_sInf, Set.biInter_range] + +/-- Subgroups of a group form a complete lattice. -/ +@[to_additive "The `AddSubgroup`s of an `AddGroup` form a complete lattice."] +instance : CompleteLattice (Subgroup G) := + { completeLatticeOfInf (Subgroup G) fun _s => + IsGLB.of_image SetLike.coe_subset_coe isGLB_biInf with + bot := ⊥ + bot_le := fun S _x hx => (mem_bot.1 hx).symm ▸ S.one_mem + top := ⊤ + le_top := fun _S x _hx => mem_top x + inf := (· ⊓ ·) + le_inf := fun _a _b _c ha hb _x hx => ⟨ha hx, hb hx⟩ + inf_le_left := fun _a _b _x => And.left + inf_le_right := fun _a _b _x => And.right } + +@[to_additive] +theorem mem_sup_left {S T : Subgroup G} : ∀ {x : G}, x ∈ S → x ∈ S ⊔ T := + have : S ≤ S ⊔ T := le_sup_left; fun h ↦ this h + +@[to_additive] +theorem mem_sup_right {S T : Subgroup G} : ∀ {x : G}, x ∈ T → x ∈ S ⊔ T := + have : T ≤ S ⊔ T := le_sup_right; fun h ↦ this h + +@[to_additive] +theorem mul_mem_sup {S T : Subgroup G} {x y : G} (hx : x ∈ S) (hy : y ∈ T) : x * y ∈ S ⊔ T := + (S ⊔ T).mul_mem (mem_sup_left hx) (mem_sup_right hy) + +@[to_additive] +theorem mem_iSup_of_mem {ι : Sort*} {S : ι → Subgroup G} (i : ι) : + ∀ {x : G}, x ∈ S i → x ∈ iSup S := + have : S i ≤ iSup S := le_iSup _ _; fun h ↦ this h + +@[to_additive] +theorem mem_sSup_of_mem {S : Set (Subgroup G)} {s : Subgroup G} (hs : s ∈ S) : + ∀ {x : G}, x ∈ s → x ∈ sSup S := + have : s ≤ sSup S := le_sSup hs; fun h ↦ this h + +@[to_additive (attr := simp)] +theorem subsingleton_iff : Subsingleton (Subgroup G) ↔ Subsingleton G := + ⟨fun _ => + ⟨fun x y => + have : ∀ i : G, i = 1 := fun i => + mem_bot.mp <| Subsingleton.elim (⊤ : Subgroup G) ⊥ ▸ mem_top i + (this x).trans (this y).symm⟩, + fun _ => ⟨fun x y => Subgroup.ext fun i => Subsingleton.elim 1 i ▸ by simp [Subgroup.one_mem]⟩⟩ + +@[to_additive (attr := simp)] +theorem nontrivial_iff : Nontrivial (Subgroup G) ↔ Nontrivial G := + not_iff_not.mp + ((not_nontrivial_iff_subsingleton.trans subsingleton_iff).trans + not_nontrivial_iff_subsingleton.symm) + +@[to_additive] +instance [Subsingleton G] : Unique (Subgroup G) := + ⟨⟨⊥⟩, fun a => @Subsingleton.elim _ (subsingleton_iff.mpr ‹_›) a _⟩ + +@[to_additive] +instance [Nontrivial G] : Nontrivial (Subgroup G) := + nontrivial_iff.mpr ‹_› + +@[to_additive] +theorem eq_top_iff' : H = ⊤ ↔ ∀ x : G, x ∈ H := + eq_top_iff.trans ⟨fun h m => h <| mem_top m, fun h m _ => h m⟩ + +/-- The `Subgroup` generated by a set. -/ +@[to_additive "The `AddSubgroup` generated by a set"] +def closure (k : Set G) : Subgroup G := + sInf { K | k ⊆ K } + +variable {k : Set G} + +@[to_additive] +theorem mem_closure {x : G} : x ∈ closure k ↔ ∀ K : Subgroup G, k ⊆ K → x ∈ K := + mem_sInf + +/-- The subgroup generated by a set includes the set. -/ +@[to_additive (attr := simp, aesop safe 20 apply (rule_sets := [SetLike])) + "The `AddSubgroup` generated by a set includes the set."] +theorem subset_closure : k ⊆ closure k := fun _ hx => mem_closure.2 fun _ hK => hK hx + +@[to_additive] +theorem not_mem_of_not_mem_closure {P : G} (hP : P ∉ closure k) : P ∉ k := fun h => + hP (subset_closure h) + +open Set + +/-- A subgroup `K` includes `closure k` if and only if it includes `k`. -/ +@[to_additive (attr := simp) + "An additive subgroup `K` includes `closure k` if and only if it includes `k`"] +theorem closure_le : closure k ≤ K ↔ k ⊆ K := + ⟨Subset.trans subset_closure, fun h => sInf_le h⟩ + +@[to_additive] +theorem closure_eq_of_le (h₁ : k ⊆ K) (h₂ : K ≤ closure k) : closure k = K := + le_antisymm ((closure_le <| K).2 h₁) h₂ + +/-- An induction principle for closure membership. If `p` holds for `1` and all elements of `k`, and +is preserved under multiplication and inverse, then `p` holds for all elements of the closure +of `k`. + +See also `Subgroup.closure_induction_left` and `Subgroup.closure_induction_right` for versions that +only require showing `p` is preserved by multiplication by elements in `k`. -/ +@[to_additive (attr := elab_as_elim) + "An induction principle for additive closure membership. If `p` + holds for `0` and all elements of `k`, and is preserved under addition and inverses, then `p` + holds for all elements of the additive closure of `k`. + + See also `AddSubgroup.closure_induction_left` and `AddSubgroup.closure_induction_left` for + versions that only require showing `p` is preserved by addition by elements in `k`."] +theorem closure_induction {p : (g : G) → g ∈ closure k → Prop} + (mem : ∀ x (hx : x ∈ k), p x (subset_closure hx)) (one : p 1 (one_mem _)) + (mul : ∀ x y hx hy, p x hx → p y hy → p (x * y) (mul_mem hx hy)) + (inv : ∀ x hx, p x hx → p x⁻¹ (inv_mem hx)) {x} (hx : x ∈ closure k) : p x hx := + let K : Subgroup G := + { carrier := { x | ∃ hx, p x hx } + mul_mem' := fun ⟨_, ha⟩ ⟨_, hb⟩ ↦ ⟨_, mul _ _ _ _ ha hb⟩ + one_mem' := ⟨_, one⟩ + inv_mem' := fun ⟨_, hb⟩ ↦ ⟨_, inv _ _ hb⟩ } + closure_le (K := K) |>.mpr (fun y hy ↦ ⟨subset_closure hy, mem y hy⟩) hx |>.elim fun _ ↦ id + +@[deprecated closure_induction (since := "2024-10-10")] +alias closure_induction' := closure_induction + +/-- An induction principle for closure membership for predicates with two arguments. -/ +@[to_additive (attr := elab_as_elim) + "An induction principle for additive closure membership, for + predicates with two arguments."] +theorem closure_induction₂ {p : (x y : G) → x ∈ closure k → y ∈ closure k → Prop} + (mem : ∀ (x) (y) (hx : x ∈ k) (hy : y ∈ k), p x y (subset_closure hx) (subset_closure hy)) + (one_left : ∀ x hx, p 1 x (one_mem _) hx) (one_right : ∀ x hx, p x 1 hx (one_mem _)) + (mul_left : ∀ x y z hx hy hz, p x z hx hz → p y z hy hz → p (x * y) z (mul_mem hx hy) hz) + (mul_right : ∀ y z x hy hz hx, p x y hx hy → p x z hx hz → p x (y * z) hx (mul_mem hy hz)) + (inv_left : ∀ x y hx hy, p x y hx hy → p x⁻¹ y (inv_mem hx) hy) + (inv_right : ∀ x y hx hy, p x y hx hy → p x y⁻¹ hx (inv_mem hy)) + {x y : G} (hx : x ∈ closure k) (hy : y ∈ closure k) : p x y hx hy := by + induction hy using closure_induction with + | mem z hz => induction hx using closure_induction with + | mem _ h => exact mem _ _ h hz + | one => exact one_left _ (subset_closure hz) + | mul _ _ _ _ h₁ h₂ => exact mul_left _ _ _ _ _ _ h₁ h₂ + | inv _ _ h => exact inv_left _ _ _ _ h + | one => exact one_right x hx + | mul _ _ _ _ h₁ h₂ => exact mul_right _ _ _ _ _ hx h₁ h₂ + | inv _ _ h => exact inv_right _ _ _ _ h + +@[to_additive (attr := simp)] +theorem closure_closure_coe_preimage {k : Set G} : closure (((↑) : closure k → G) ⁻¹' k) = ⊤ := + eq_top_iff.2 fun x _ ↦ Subtype.recOn x fun _ hx' ↦ + closure_induction (fun _ h ↦ subset_closure h) (one_mem _) (fun _ _ _ _ ↦ mul_mem) + (fun _ _ ↦ inv_mem) hx' + +variable (G) + +/-- `closure` forms a Galois insertion with the coercion to set. -/ +@[to_additive "`closure` forms a Galois insertion with the coercion to set."] +protected def gi : GaloisInsertion (@closure G _) (↑) where + choice s _ := closure s + gc s t := @closure_le _ _ t s + le_l_u _s := subset_closure + choice_eq _s _h := rfl + +variable {G} + +/-- Subgroup closure of a set is monotone in its argument: if `h ⊆ k`, +then `closure h ≤ closure k`. -/ +@[to_additive + "Additive subgroup closure of a set is monotone in its argument: if `h ⊆ k`, + then `closure h ≤ closure k`"] +theorem closure_mono ⦃h k : Set G⦄ (h' : h ⊆ k) : closure h ≤ closure k := + (Subgroup.gi G).gc.monotone_l h' + +/-- Closure of a subgroup `K` equals `K`. -/ +@[to_additive (attr := simp) "Additive closure of an additive subgroup `K` equals `K`"] +theorem closure_eq : closure (K : Set G) = K := + (Subgroup.gi G).l_u_eq K + +@[to_additive (attr := simp)] +theorem closure_empty : closure (∅ : Set G) = ⊥ := + (Subgroup.gi G).gc.l_bot + +@[to_additive (attr := simp)] +theorem closure_univ : closure (univ : Set G) = ⊤ := + @coe_top G _ ▸ closure_eq ⊤ + +@[to_additive] +theorem closure_union (s t : Set G) : closure (s ∪ t) = closure s ⊔ closure t := + (Subgroup.gi G).gc.l_sup + +@[to_additive] +theorem sup_eq_closure (H H' : Subgroup G) : H ⊔ H' = closure ((H : Set G) ∪ (H' : Set G)) := by + simp_rw [closure_union, closure_eq] + +@[to_additive] +theorem closure_iUnion {ι} (s : ι → Set G) : closure (⋃ i, s i) = ⨆ i, closure (s i) := + (Subgroup.gi G).gc.l_iSup + +@[to_additive (attr := simp)] +theorem closure_eq_bot_iff : closure k = ⊥ ↔ k ⊆ {1} := le_bot_iff.symm.trans <| closure_le _ + +@[to_additive] +theorem iSup_eq_closure {ι : Sort*} (p : ι → Subgroup G) : + ⨆ i, p i = closure (⋃ i, (p i : Set G)) := by simp_rw [closure_iUnion, closure_eq] + +/-- The subgroup generated by an element of a group equals the set of integer number powers of + the element. -/ +@[to_additive + "The `AddSubgroup` generated by an element of an `AddGroup` equals the set of + natural number multiples of the element."] +theorem mem_closure_singleton {x y : G} : y ∈ closure ({x} : Set G) ↔ ∃ n : ℤ, x ^ n = y := by + refine + ⟨fun hy => closure_induction ?_ ?_ ?_ ?_ hy, fun ⟨n, hn⟩ => + hn ▸ zpow_mem (subset_closure <| mem_singleton x) n⟩ + · intro y hy + rw [eq_of_mem_singleton hy] + exact ⟨1, zpow_one x⟩ + · exact ⟨0, zpow_zero x⟩ + · rintro _ _ _ _ ⟨n, rfl⟩ ⟨m, rfl⟩ + exact ⟨n + m, zpow_add x n m⟩ + rintro _ _ ⟨n, rfl⟩ + exact ⟨-n, zpow_neg x n⟩ + +@[to_additive] +theorem closure_singleton_one : closure ({1} : Set G) = ⊥ := by + simp [eq_bot_iff_forall, mem_closure_singleton] + +@[to_additive (attr := simp)] +lemma mem_closure_singleton_self (x : G) : x ∈ closure ({x} : Set G) := by + simpa [-subset_closure] using subset_closure (k := {x}) + +@[to_additive] +theorem le_closure_toSubmonoid (S : Set G) : Submonoid.closure S ≤ (closure S).toSubmonoid := + Submonoid.closure_le.2 subset_closure + +@[to_additive] +theorem closure_eq_top_of_mclosure_eq_top {S : Set G} (h : Submonoid.closure S = ⊤) : + closure S = ⊤ := + (eq_top_iff' _).2 fun _ => le_closure_toSubmonoid _ <| h.symm ▸ trivial + +@[to_additive] +theorem mem_iSup_of_directed {ι} [hι : Nonempty ι] {K : ι → Subgroup G} (hK : Directed (· ≤ ·) K) + {x : G} : x ∈ (iSup K : Subgroup G) ↔ ∃ i, x ∈ K i := by + refine ⟨?_, fun ⟨i, hi⟩ ↦ le_iSup K i hi⟩ + suffices x ∈ closure (⋃ i, (K i : Set G)) → ∃ i, x ∈ K i by + simpa only [closure_iUnion, closure_eq (K _)] using this + refine fun hx ↦ closure_induction (fun _ ↦ mem_iUnion.1) ?_ ?_ ?_ hx + · exact hι.elim fun i ↦ ⟨i, (K i).one_mem⟩ + · rintro x y _ _ ⟨i, hi⟩ ⟨j, hj⟩ + rcases hK i j with ⟨k, hki, hkj⟩ + exact ⟨k, mul_mem (hki hi) (hkj hj)⟩ + · rintro _ _ ⟨i, hi⟩ + exact ⟨i, inv_mem hi⟩ + +@[to_additive] +theorem coe_iSup_of_directed {ι} [Nonempty ι] {S : ι → Subgroup G} (hS : Directed (· ≤ ·) S) : + ((⨆ i, S i : Subgroup G) : Set G) = ⋃ i, S i := + Set.ext fun x ↦ by simp [mem_iSup_of_directed hS] + +@[to_additive] +theorem mem_sSup_of_directedOn {K : Set (Subgroup G)} (Kne : K.Nonempty) (hK : DirectedOn (· ≤ ·) K) + {x : G} : x ∈ sSup K ↔ ∃ s ∈ K, x ∈ s := by + haveI : Nonempty K := Kne.to_subtype + simp only [sSup_eq_iSup', mem_iSup_of_directed hK.directed_val, SetCoe.exists, Subtype.coe_mk, + exists_prop] + +variable {C : Type*} [CommGroup C] {s t : Subgroup C} {x : C} + +@[to_additive] +theorem mem_sup : x ∈ s ⊔ t ↔ ∃ y ∈ s, ∃ z ∈ t, y * z = x := + ⟨fun h => by + rw [sup_eq_closure] at h + refine Subgroup.closure_induction ?_ ?_ ?_ ?_ h + · rintro y (h | h) + · exact ⟨y, h, 1, t.one_mem, by simp⟩ + · exact ⟨1, s.one_mem, y, h, by simp⟩ + · exact ⟨1, s.one_mem, 1, ⟨t.one_mem, mul_one 1⟩⟩ + · rintro _ _ _ _ ⟨y₁, hy₁, z₁, hz₁, rfl⟩ ⟨y₂, hy₂, z₂, hz₂, rfl⟩ + exact ⟨_, mul_mem hy₁ hy₂, _, mul_mem hz₁ hz₂, by simp [mul_assoc, mul_left_comm]⟩ + · rintro _ _ ⟨y, hy, z, hz, rfl⟩ + exact ⟨_, inv_mem hy, _, inv_mem hz, mul_comm z y ▸ (mul_inv_rev z y).symm⟩, by + rintro ⟨y, hy, z, hz, rfl⟩; exact mul_mem_sup hy hz⟩ + +@[to_additive] +theorem mem_sup' : x ∈ s ⊔ t ↔ ∃ (y : s) (z : t), (y : C) * z = x := + mem_sup.trans <| by simp only [SetLike.exists, coe_mk, exists_prop] + +@[to_additive] +theorem mem_closure_pair {x y z : C} : + z ∈ closure ({x, y} : Set C) ↔ ∃ m n : ℤ, x ^ m * y ^ n = z := by + rw [← Set.singleton_union, Subgroup.closure_union, mem_sup] + simp_rw [mem_closure_singleton, exists_exists_eq_and] + +@[to_additive] +theorem disjoint_def {H₁ H₂ : Subgroup G} : Disjoint H₁ H₂ ↔ ∀ {x : G}, x ∈ H₁ → x ∈ H₂ → x = 1 := + disjoint_iff_inf_le.trans <| by simp only [Disjoint, SetLike.le_def, mem_inf, mem_bot, and_imp] + +@[to_additive] +theorem disjoint_def' {H₁ H₂ : Subgroup G} : + Disjoint H₁ H₂ ↔ ∀ {x y : G}, x ∈ H₁ → y ∈ H₂ → x = y → x = 1 := + disjoint_def.trans ⟨fun h _x _y hx hy hxy ↦ h hx <| hxy.symm ▸ hy, fun h _x hx hx' ↦ h hx hx' rfl⟩ + +@[to_additive] +theorem disjoint_iff_mul_eq_one {H₁ H₂ : Subgroup G} : + Disjoint H₁ H₂ ↔ ∀ {x y : G}, x ∈ H₁ → y ∈ H₂ → x * y = 1 → x = 1 ∧ y = 1 := + disjoint_def'.trans + ⟨fun h x y hx hy hxy => + let hx1 : x = 1 := h hx (H₂.inv_mem hy) (eq_inv_iff_mul_eq_one.mpr hxy) + ⟨hx1, by simpa [hx1] using hxy⟩, + fun h _ _ hx hy hxy => (h hx (H₂.inv_mem hy) (mul_inv_eq_one.mpr hxy)).1⟩ + +@[to_additive] +theorem mul_injective_of_disjoint {H₁ H₂ : Subgroup G} (h : Disjoint H₁ H₂) : + Function.Injective (fun g => g.1 * g.2 : H₁ × H₂ → G) := by + intro x y hxy + rw [← inv_mul_eq_iff_eq_mul, ← mul_assoc, ← mul_inv_eq_one, mul_assoc] at hxy + replace hxy := disjoint_iff_mul_eq_one.mp h (y.1⁻¹ * x.1).prop (x.2 * y.2⁻¹).prop hxy + rwa [coe_mul, coe_mul, coe_inv, coe_inv, inv_mul_eq_one, mul_inv_eq_one, ← Subtype.ext_iff, ← + Subtype.ext_iff, eq_comm, ← Prod.ext_iff] at hxy + +end Subgroup diff --git a/Mathlib/Algebra/Group/Subgroup/Map.lean b/Mathlib/Algebra/Group/Subgroup/Map.lean new file mode 100644 index 0000000000000..84450e745384e --- /dev/null +++ b/Mathlib/Algebra/Group/Subgroup/Map.lean @@ -0,0 +1,527 @@ +/- +Copyright (c) 2020 Kexing Ying. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kexing Ying +-/ +import Mathlib.Algebra.Group.Subgroup.Lattice + +/-! +# `map` and `comap` for subgroups + +We prove results about images and preimages of subgroups under group homomorphisms. The bundled +subgroups use bundled monoid homomorphisms. + +Special thanks goes to Amelia Livingston and Yury Kudryashov for their help and inspiration. + +## Main definitions + +Notation used here: + +- `G N` are `Group`s + +- `H` is a `Subgroup` of `G` + +- `x` is an element of type `G` or type `A` + +- `f g : N →* G` are group homomorphisms + +- `s k` are sets of elements of type `G` + +Definitions in the file: + +* `Subgroup.comap H f` : the preimage of a subgroup `H` along the group homomorphism `f` is also a + subgroup + +* `Subgroup.map f H` : the image of a subgroup `H` along the group homomorphism `f` is also a + subgroup + +## Implementation notes + +Subgroup inclusion is denoted `≤` rather than `⊆`, although `∈` is defined as +membership of a subgroup's underlying set. + +## Tags +subgroup, subgroups +-/ + +assert_not_exists OrderedAddCommMonoid +assert_not_exists Multiset +assert_not_exists Ring + +open Function +open scoped Int + +variable {G G' G'' : Type*} [Group G] [Group G'] [Group G''] +variable {A : Type*} [AddGroup A] + +namespace Subgroup + +variable (H K : Subgroup G) {k : Set G} + +open Set + +variable {N : Type*} [Group N] {P : Type*} [Group P] + +/-- The preimage of a subgroup along a monoid homomorphism is a subgroup. -/ +@[to_additive + "The preimage of an `AddSubgroup` along an `AddMonoid` homomorphism + is an `AddSubgroup`."] +def comap {N : Type*} [Group N] (f : G →* N) (H : Subgroup N) : Subgroup G := + { H.toSubmonoid.comap f with + carrier := f ⁻¹' H + inv_mem' := fun {a} ha => show f a⁻¹ ∈ H by rw [f.map_inv]; exact H.inv_mem ha } + +@[to_additive (attr := simp)] +theorem coe_comap (K : Subgroup N) (f : G →* N) : (K.comap f : Set G) = f ⁻¹' K := + rfl + +@[to_additive (attr := simp)] +theorem mem_comap {K : Subgroup N} {f : G →* N} {x : G} : x ∈ K.comap f ↔ f x ∈ K := + Iff.rfl + +@[to_additive] +theorem comap_mono {f : G →* N} {K K' : Subgroup N} : K ≤ K' → comap f K ≤ comap f K' := + preimage_mono + +@[to_additive] +theorem comap_comap (K : Subgroup P) (g : N →* P) (f : G →* N) : + (K.comap g).comap f = K.comap (g.comp f) := + rfl + +@[to_additive (attr := simp)] +theorem comap_id (K : Subgroup N) : K.comap (MonoidHom.id _) = K := by + ext + rfl + +@[simp] +theorem toAddSubgroup_comap {G₂ : Type*} [Group G₂] (f : G →* G₂) (s : Subgroup G₂) : + s.toAddSubgroup.comap (MonoidHom.toAdditive f) = Subgroup.toAddSubgroup (s.comap f) := rfl + +@[simp] +theorem _root_.AddSubgroup.toSubgroup_comap {A A₂ : Type*} [AddGroup A] [AddGroup A₂] + (f : A →+ A₂) (s : AddSubgroup A₂) : + s.toSubgroup.comap (AddMonoidHom.toMultiplicative f) = AddSubgroup.toSubgroup (s.comap f) := rfl + +/-- The image of a subgroup along a monoid homomorphism is a subgroup. -/ +@[to_additive + "The image of an `AddSubgroup` along an `AddMonoid` homomorphism + is an `AddSubgroup`."] +def map (f : G →* N) (H : Subgroup G) : Subgroup N := + { H.toSubmonoid.map f with + carrier := f '' H + inv_mem' := by + rintro _ ⟨x, hx, rfl⟩ + exact ⟨x⁻¹, H.inv_mem hx, f.map_inv x⟩ } + +@[to_additive (attr := simp)] +theorem coe_map (f : G →* N) (K : Subgroup G) : (K.map f : Set N) = f '' K := + rfl + +@[to_additive (attr := simp)] +theorem mem_map {f : G →* N} {K : Subgroup G} {y : N} : y ∈ K.map f ↔ ∃ x ∈ K, f x = y := Iff.rfl + +@[to_additive] +theorem mem_map_of_mem (f : G →* N) {K : Subgroup G} {x : G} (hx : x ∈ K) : f x ∈ K.map f := + mem_image_of_mem f hx + +@[to_additive] +theorem apply_coe_mem_map (f : G →* N) (K : Subgroup G) (x : K) : f x ∈ K.map f := + mem_map_of_mem f x.prop + +@[to_additive] +theorem map_mono {f : G →* N} {K K' : Subgroup G} : K ≤ K' → map f K ≤ map f K' := + image_subset _ + +@[to_additive (attr := simp)] +theorem map_id : K.map (MonoidHom.id G) = K := + SetLike.coe_injective <| image_id _ + +@[to_additive] +theorem map_map (g : N →* P) (f : G →* N) : (K.map f).map g = K.map (g.comp f) := + SetLike.coe_injective <| image_image _ _ _ + +@[to_additive (attr := simp)] +theorem map_one_eq_bot : K.map (1 : G →* N) = ⊥ := + eq_bot_iff.mpr <| by + rintro x ⟨y, _, rfl⟩ + simp + +@[to_additive] +theorem mem_map_equiv {f : G ≃* N} {K : Subgroup G} {x : N} : + x ∈ K.map f.toMonoidHom ↔ f.symm x ∈ K := by + erw [@Set.mem_image_equiv _ _ (↑K) f.toEquiv x]; rfl + +-- The simpNF linter says that the LHS can be simplified via `Subgroup.mem_map`. +-- However this is a higher priority lemma. +-- https://github.com/leanprover/std4/issues/207 +@[to_additive (attr := simp 1100, nolint simpNF)] +theorem mem_map_iff_mem {f : G →* N} (hf : Function.Injective f) {K : Subgroup G} {x : G} : + f x ∈ K.map f ↔ x ∈ K := + hf.mem_set_image + +@[to_additive] +theorem map_equiv_eq_comap_symm' (f : G ≃* N) (K : Subgroup G) : + K.map f.toMonoidHom = K.comap f.symm.toMonoidHom := + SetLike.coe_injective (f.toEquiv.image_eq_preimage K) + +@[to_additive] +theorem map_equiv_eq_comap_symm (f : G ≃* N) (K : Subgroup G) : + K.map f = K.comap (G := N) f.symm := + map_equiv_eq_comap_symm' _ _ + +@[to_additive] +theorem comap_equiv_eq_map_symm (f : N ≃* G) (K : Subgroup G) : + K.comap (G := N) f = K.map f.symm := + (map_equiv_eq_comap_symm f.symm K).symm + +@[to_additive] +theorem comap_equiv_eq_map_symm' (f : N ≃* G) (K : Subgroup G) : + K.comap f.toMonoidHom = K.map f.symm.toMonoidHom := + (map_equiv_eq_comap_symm f.symm K).symm + +@[to_additive] +theorem map_symm_eq_iff_map_eq {H : Subgroup N} {e : G ≃* N} : + H.map ↑e.symm = K ↔ K.map ↑e = H := by + constructor <;> rintro rfl + · rw [map_map, ← MulEquiv.coe_monoidHom_trans, MulEquiv.symm_trans_self, + MulEquiv.coe_monoidHom_refl, map_id] + · rw [map_map, ← MulEquiv.coe_monoidHom_trans, MulEquiv.self_trans_symm, + MulEquiv.coe_monoidHom_refl, map_id] + +@[to_additive] +theorem map_le_iff_le_comap {f : G →* N} {K : Subgroup G} {H : Subgroup N} : + K.map f ≤ H ↔ K ≤ H.comap f := + image_subset_iff + +@[to_additive] +theorem gc_map_comap (f : G →* N) : GaloisConnection (map f) (comap f) := fun _ _ => + map_le_iff_le_comap + +@[to_additive] +theorem map_sup (H K : Subgroup G) (f : G →* N) : (H ⊔ K).map f = H.map f ⊔ K.map f := + (gc_map_comap f).l_sup + +@[to_additive] +theorem map_iSup {ι : Sort*} (f : G →* N) (s : ι → Subgroup G) : + (iSup s).map f = ⨆ i, (s i).map f := + (gc_map_comap f).l_iSup + +@[to_additive] +theorem map_inf (H K : Subgroup G) (f : G →* N) (hf : Function.Injective f) : + (H ⊓ K).map f = H.map f ⊓ K.map f := SetLike.coe_injective (Set.image_inter hf) + +@[to_additive] +theorem map_iInf {ι : Sort*} [Nonempty ι] (f : G →* N) (hf : Function.Injective f) + (s : ι → Subgroup G) : (iInf s).map f = ⨅ i, (s i).map f := by + apply SetLike.coe_injective + simpa using (Set.injOn_of_injective hf).image_iInter_eq (s := SetLike.coe ∘ s) + +@[to_additive] +theorem comap_sup_comap_le (H K : Subgroup N) (f : G →* N) : + comap f H ⊔ comap f K ≤ comap f (H ⊔ K) := + Monotone.le_map_sup (fun _ _ => comap_mono) H K + +@[to_additive] +theorem iSup_comap_le {ι : Sort*} (f : G →* N) (s : ι → Subgroup N) : + ⨆ i, (s i).comap f ≤ (iSup s).comap f := + Monotone.le_map_iSup fun _ _ => comap_mono + +@[to_additive] +theorem comap_inf (H K : Subgroup N) (f : G →* N) : (H ⊓ K).comap f = H.comap f ⊓ K.comap f := + (gc_map_comap f).u_inf + +@[to_additive] +theorem comap_iInf {ι : Sort*} (f : G →* N) (s : ι → Subgroup N) : + (iInf s).comap f = ⨅ i, (s i).comap f := + (gc_map_comap f).u_iInf + +@[to_additive] +theorem map_inf_le (H K : Subgroup G) (f : G →* N) : map f (H ⊓ K) ≤ map f H ⊓ map f K := + le_inf (map_mono inf_le_left) (map_mono inf_le_right) + +@[to_additive] +theorem map_inf_eq (H K : Subgroup G) (f : G →* N) (hf : Function.Injective f) : + map f (H ⊓ K) = map f H ⊓ map f K := by + rw [← SetLike.coe_set_eq] + simp [Set.image_inter hf] + +@[to_additive (attr := simp)] +theorem map_bot (f : G →* N) : (⊥ : Subgroup G).map f = ⊥ := + (gc_map_comap f).l_bot + +@[to_additive (attr := simp)] +theorem map_top_of_surjective (f : G →* N) (h : Function.Surjective f) : Subgroup.map f ⊤ = ⊤ := by + rw [eq_top_iff] + intro x _ + obtain ⟨y, hy⟩ := h x + exact ⟨y, trivial, hy⟩ + +@[to_additive (attr := simp)] +theorem comap_top (f : G →* N) : (⊤ : Subgroup N).comap f = ⊤ := + (gc_map_comap f).u_top + +/-- For any subgroups `H` and `K`, view `H ⊓ K` as a subgroup of `K`. -/ +@[to_additive "For any subgroups `H` and `K`, view `H ⊓ K` as a subgroup of `K`."] +def subgroupOf (H K : Subgroup G) : Subgroup K := + H.comap K.subtype + +/-- If `H ≤ K`, then `H` as a subgroup of `K` is isomorphic to `H`. -/ +@[to_additive (attr := simps) "If `H ≤ K`, then `H` as a subgroup of `K` is isomorphic to `H`."] +def subgroupOfEquivOfLe {G : Type*} [Group G] {H K : Subgroup G} (h : H ≤ K) : + H.subgroupOf K ≃* H where + toFun g := ⟨g.1, g.2⟩ + invFun g := ⟨⟨g.1, h g.2⟩, g.2⟩ + left_inv _g := Subtype.ext (Subtype.ext rfl) + right_inv _g := Subtype.ext rfl + map_mul' _g _h := rfl + +@[to_additive (attr := simp)] +theorem comap_subtype (H K : Subgroup G) : H.comap K.subtype = H.subgroupOf K := + rfl + +@[to_additive (attr := simp)] +theorem comap_inclusion_subgroupOf {K₁ K₂ : Subgroup G} (h : K₁ ≤ K₂) (H : Subgroup G) : + (H.subgroupOf K₂).comap (inclusion h) = H.subgroupOf K₁ := + rfl + +@[to_additive] +theorem coe_subgroupOf (H K : Subgroup G) : (H.subgroupOf K : Set K) = K.subtype ⁻¹' H := + rfl + +@[to_additive] +theorem mem_subgroupOf {H K : Subgroup G} {h : K} : h ∈ H.subgroupOf K ↔ (h : G) ∈ H := + Iff.rfl + +-- TODO(kmill): use `K ⊓ H` order for RHS to match `Subtype.image_preimage_coe` +@[to_additive (attr := simp)] +theorem subgroupOf_map_subtype (H K : Subgroup G) : (H.subgroupOf K).map K.subtype = H ⊓ K := + SetLike.ext' <| by refine Subtype.image_preimage_coe _ _ |>.trans ?_; apply Set.inter_comm + +@[to_additive (attr := simp)] +theorem bot_subgroupOf : (⊥ : Subgroup G).subgroupOf H = ⊥ := + Eq.symm (Subgroup.ext fun _g => Subtype.ext_iff) + +@[to_additive (attr := simp)] +theorem top_subgroupOf : (⊤ : Subgroup G).subgroupOf H = ⊤ := + rfl + +@[to_additive] +theorem subgroupOf_bot_eq_bot : H.subgroupOf ⊥ = ⊥ := + Subsingleton.elim _ _ + +@[to_additive] +theorem subgroupOf_bot_eq_top : H.subgroupOf ⊥ = ⊤ := + Subsingleton.elim _ _ + +@[to_additive (attr := simp)] +theorem subgroupOf_self : H.subgroupOf H = ⊤ := + top_unique fun g _hg => g.2 + +@[to_additive (attr := simp)] +theorem subgroupOf_inj {H₁ H₂ K : Subgroup G} : + H₁.subgroupOf K = H₂.subgroupOf K ↔ H₁ ⊓ K = H₂ ⊓ K := by + simpa only [SetLike.ext_iff, mem_inf, mem_subgroupOf, and_congr_left_iff] using Subtype.forall + +@[to_additive (attr := simp)] +theorem inf_subgroupOf_right (H K : Subgroup G) : (H ⊓ K).subgroupOf K = H.subgroupOf K := + subgroupOf_inj.2 (inf_right_idem _ _) + +@[to_additive (attr := simp)] +theorem inf_subgroupOf_left (H K : Subgroup G) : (K ⊓ H).subgroupOf K = H.subgroupOf K := by + rw [inf_comm, inf_subgroupOf_right] + +@[to_additive (attr := simp)] +theorem subgroupOf_eq_bot {H K : Subgroup G} : H.subgroupOf K = ⊥ ↔ Disjoint H K := by + rw [disjoint_iff, ← bot_subgroupOf, subgroupOf_inj, bot_inf_eq] + +@[to_additive (attr := simp)] +theorem subgroupOf_eq_top {H K : Subgroup G} : H.subgroupOf K = ⊤ ↔ K ≤ H := by + rw [← top_subgroupOf, subgroupOf_inj, top_inf_eq, inf_eq_right] + +variable (H : Subgroup G) + +@[to_additive] +instance map_isCommutative (f : G →* G') [H.IsCommutative] : (H.map f).IsCommutative := + ⟨⟨by + rintro ⟨-, a, ha, rfl⟩ ⟨-, b, hb, rfl⟩ + rw [Subtype.ext_iff, coe_mul, coe_mul, Subtype.coe_mk, Subtype.coe_mk, ← map_mul, ← map_mul] + exact congr_arg f (Subtype.ext_iff.mp (mul_comm (⟨a, ha⟩ : H) ⟨b, hb⟩))⟩⟩ + +@[to_additive] +theorem comap_injective_isCommutative {f : G' →* G} (hf : Injective f) [H.IsCommutative] : + (H.comap f).IsCommutative := + ⟨⟨fun a b => + Subtype.ext + (by + have := mul_comm (⟨f a, a.2⟩ : H) (⟨f b, b.2⟩ : H) + rwa [Subtype.ext_iff, coe_mul, coe_mul, coe_mk, coe_mk, ← map_mul, ← map_mul, + hf.eq_iff] at this)⟩⟩ + +@[to_additive] +instance subgroupOf_isCommutative [H.IsCommutative] : (H.subgroupOf K).IsCommutative := + H.comap_injective_isCommutative Subtype.coe_injective + +end Subgroup + +namespace MulEquiv +variable {H : Type*} [Group H] + +/-- +An isomorphism of groups gives an order isomorphism between the lattices of subgroups, +defined by sending subgroups to their inverse images. + +See also `MulEquiv.mapSubgroup` which maps subgroups to their forward images. +-/ +@[simps] +def comapSubgroup (f : G ≃* H) : Subgroup H ≃o Subgroup G where + toFun := Subgroup.comap f + invFun := Subgroup.comap f.symm + left_inv sg := by simp [Subgroup.comap_comap] + right_inv sh := by simp [Subgroup.comap_comap] + map_rel_iff' {sg1 sg2} := + ⟨fun h => by simpa [Subgroup.comap_comap] using + Subgroup.comap_mono (f := (f.symm : H →* G)) h, Subgroup.comap_mono⟩ + +/-- +An isomorphism of groups gives an order isomorphism between the lattices of subgroups, +defined by sending subgroups to their forward images. + +See also `MulEquiv.comapSubgroup` which maps subgroups to their inverse images. +-/ +@[simps] +def mapSubgroup {H : Type*} [Group H] (f : G ≃* H) : Subgroup G ≃o Subgroup H where + toFun := Subgroup.map f + invFun := Subgroup.map f.symm + left_inv sg := by simp [Subgroup.map_map] + right_inv sh := by simp [Subgroup.map_map] + map_rel_iff' {sg1 sg2} := + ⟨fun h => by simpa [Subgroup.map_map] using + Subgroup.map_mono (f := (f.symm : H →* G)) h, Subgroup.map_mono⟩ + +end MulEquiv + +namespace Subgroup + +open MonoidHom + +variable {N : Type*} [Group N] (f : G →* N) + +@[to_additive] +theorem map_comap_le (H : Subgroup N) : map f (comap f H) ≤ H := + (gc_map_comap f).l_u_le _ + +@[to_additive] +theorem le_comap_map (H : Subgroup G) : H ≤ comap f (map f H) := + (gc_map_comap f).le_u_l _ + +@[to_additive] +theorem map_eq_comap_of_inverse {f : G →* N} {g : N →* G} (hl : Function.LeftInverse g f) + (hr : Function.RightInverse g f) (H : Subgroup G) : map f H = comap g H := + SetLike.ext' <| by rw [coe_map, coe_comap, Set.image_eq_preimage_of_inverse hl hr] + +/-- A subgroup is isomorphic to its image under an injective function. If you have an isomorphism, +use `MulEquiv.subgroupMap` for better definitional equalities. -/ +@[to_additive + "An additive subgroup is isomorphic to its image under an injective function. If you + have an isomorphism, use `AddEquiv.addSubgroupMap` for better definitional equalities."] +noncomputable def equivMapOfInjective (H : Subgroup G) (f : G →* N) (hf : Function.Injective f) : + H ≃* H.map f := + { Equiv.Set.image f H hf with map_mul' := fun _ _ => Subtype.ext (f.map_mul _ _) } + +@[to_additive (attr := simp)] +theorem coe_equivMapOfInjective_apply (H : Subgroup G) (f : G →* N) (hf : Function.Injective f) + (h : H) : (equivMapOfInjective H f hf h : N) = f h := + rfl + +end Subgroup + +variable {N : Type*} [Group N] + +namespace MonoidHom + +/-- The `MonoidHom` from the preimage of a subgroup to itself. -/ +@[to_additive (attr := simps!) "the `AddMonoidHom` from the preimage of an +additive subgroup to itself."] +def subgroupComap (f : G →* G') (H' : Subgroup G') : H'.comap f →* H' := + f.submonoidComap H'.toSubmonoid + +/-- The `MonoidHom` from a subgroup to its image. -/ +@[to_additive (attr := simps!) "the `AddMonoidHom` from an additive subgroup to its image"] +def subgroupMap (f : G →* G') (H : Subgroup G) : H →* H.map f := + f.submonoidMap H.toSubmonoid + +@[to_additive] +theorem subgroupMap_surjective (f : G →* G') (H : Subgroup G) : + Function.Surjective (f.subgroupMap H) := + f.submonoidMap_surjective H.toSubmonoid + +end MonoidHom + +namespace MulEquiv + +variable {H K : Subgroup G} + +/-- Makes the identity isomorphism from a proof two subgroups of a multiplicative + group are equal. -/ +@[to_additive + "Makes the identity additive isomorphism from a proof + two subgroups of an additive group are equal."] +def subgroupCongr (h : H = K) : H ≃* K := + { Equiv.setCongr <| congr_arg _ h with map_mul' := fun _ _ => rfl } + +@[to_additive (attr := simp)] +lemma subgroupCongr_apply (h : H = K) (x) : + (MulEquiv.subgroupCongr h x : G) = x := rfl + +@[to_additive (attr := simp)] +lemma subgroupCongr_symm_apply (h : H = K) (x) : + ((MulEquiv.subgroupCongr h).symm x : G) = x := rfl + +/-- A subgroup is isomorphic to its image under an isomorphism. If you only have an injective map, +use `Subgroup.equiv_map_of_injective`. -/ +@[to_additive + "An additive subgroup is isomorphic to its image under an isomorphism. If you only + have an injective map, use `AddSubgroup.equiv_map_of_injective`."] +def subgroupMap (e : G ≃* G') (H : Subgroup G) : H ≃* H.map (e : G →* G') := + MulEquiv.submonoidMap (e : G ≃* G') H.toSubmonoid + +@[to_additive (attr := simp)] +theorem coe_subgroupMap_apply (e : G ≃* G') (H : Subgroup G) (g : H) : + ((subgroupMap e H g : H.map (e : G →* G')) : G') = e g := + rfl + +@[to_additive (attr := simp)] +theorem subgroupMap_symm_apply (e : G ≃* G') (H : Subgroup G) (g : H.map (e : G →* G')) : + (e.subgroupMap H).symm g = ⟨e.symm g, SetLike.mem_coe.1 <| Set.mem_image_equiv.1 g.2⟩ := + rfl + +end MulEquiv + +namespace MonoidHom + +open Subgroup + +@[to_additive] +theorem closure_preimage_le (f : G →* N) (s : Set N) : closure (f ⁻¹' s) ≤ (closure s).comap f := + (closure_le _).2 fun x hx => by rw [SetLike.mem_coe, mem_comap]; exact subset_closure hx + +/-- The image under a monoid homomorphism of the subgroup generated by a set equals the subgroup +generated by the image of the set. -/ +@[to_additive + "The image under an `AddMonoid` hom of the `AddSubgroup` generated by a set equals + the `AddSubgroup` generated by the image of the set."] +theorem map_closure (f : G →* N) (s : Set G) : (closure s).map f = closure (f '' s) := + Set.image_preimage.l_comm_of_u_comm (gc_map_comap f) (Subgroup.gi N).gc (Subgroup.gi G).gc + fun _ ↦ rfl + +end MonoidHom + +namespace Subgroup + +@[to_additive (attr := simp)] +theorem equivMapOfInjective_coe_mulEquiv (H : Subgroup G) (e : G ≃* G') : + H.equivMapOfInjective (e : G →* G') (EquivLike.injective e) = e.subgroupMap H := by + ext + rfl + +end Subgroup diff --git a/Mathlib/Algebra/Group/Subgroup/ZPowers/Basic.lean b/Mathlib/Algebra/Group/Subgroup/ZPowers/Basic.lean index cf8febc23281b..357b32013f358 100644 --- a/Mathlib/Algebra/Group/Subgroup/ZPowers/Basic.lean +++ b/Mathlib/Algebra/Group/Subgroup/ZPowers/Basic.lean @@ -3,7 +3,8 @@ Copyright (c) 2020 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import Mathlib.Algebra.Group.Subgroup.Basic +import Mathlib.Algebra.Group.Int +import Mathlib.Algebra.Group.Subgroup.Map /-! # Subgroups generated by an element diff --git a/Mathlib/Algebra/Group/Submonoid/Units.lean b/Mathlib/Algebra/Group/Submonoid/Units.lean index 96ce6d155d1e4..111adf47e20ff 100644 --- a/Mathlib/Algebra/Group/Submonoid/Units.lean +++ b/Mathlib/Algebra/Group/Submonoid/Units.lean @@ -5,7 +5,7 @@ Authors: Wrenna Robson -/ import Mathlib.Algebra.Group.Submonoid.Operations import Mathlib.Algebra.Group.Submonoid.Pointwise -import Mathlib.Algebra.Group.Subgroup.Basic +import Mathlib.Algebra.Group.Subgroup.Lattice /-! diff --git a/Mathlib/Algebra/Module/Submodule/Ker.lean b/Mathlib/Algebra/Module/Submodule/Ker.lean index b291dd847675d..dfa7ba1ee50ad 100644 --- a/Mathlib/Algebra/Module/Submodule/Ker.lean +++ b/Mathlib/Algebra/Module/Submodule/Ker.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Kevin Buzzard, Yury Kudryashov, Frédéric Dupuis, Heather Macbeth -/ +import Mathlib.Algebra.Group.Subgroup.Ker import Mathlib.Algebra.Module.Submodule.Map /-! diff --git a/Mathlib/Algebra/Module/Submodule/Lattice.lean b/Mathlib/Algebra/Module/Submodule/Lattice.lean index 8a537d9a01ba5..76744e1fa2be2 100644 --- a/Mathlib/Algebra/Module/Submodule/Lattice.lean +++ b/Mathlib/Algebra/Module/Submodule/Lattice.lean @@ -3,7 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Kevin Buzzard, Yury Kudryashov -/ -import Mathlib.Algebra.Group.Subgroup.Basic +import Mathlib.Algebra.Group.Subgroup.Lattice import Mathlib.Algebra.Group.Submonoid.Membership import Mathlib.Algebra.Module.Submodule.Defs import Mathlib.Algebra.Module.Equiv.Defs diff --git a/Mathlib/Algebra/Module/Submodule/Map.lean b/Mathlib/Algebra/Module/Submodule/Map.lean index bf9a103f9a34a..ac43875d2bd8a 100644 --- a/Mathlib/Algebra/Module/Submodule/Map.lean +++ b/Mathlib/Algebra/Module/Submodule/Map.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Kevin Buzzard, Yury Kudryashov, Frédéric Dupuis, Heather Macbeth -/ - +import Mathlib.Algebra.Group.Subgroup.Map import Mathlib.Algebra.Module.Submodule.Basic import Mathlib.Algebra.Module.Submodule.Lattice import Mathlib.Algebra.Module.Submodule.LinearMap diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index 4242b77e0dfd6..91cdbe19d79a5 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Johannes Hölzl, Yaël Dillies -/ import Mathlib.Algebra.CharP.Defs -import Mathlib.Algebra.Group.Subgroup.Basic +import Mathlib.Algebra.Group.Subgroup.Ker import Mathlib.Analysis.Normed.Group.Seminorm import Mathlib.Topology.Metrizable.Uniformity import Mathlib.Topology.Sequences diff --git a/Mathlib/Combinatorics/Derangements/Basic.lean b/Mathlib/Combinatorics/Derangements/Basic.lean index 9620588f45486..ac4e2f3b84276 100644 --- a/Mathlib/Combinatorics/Derangements/Basic.lean +++ b/Mathlib/Combinatorics/Derangements/Basic.lean @@ -7,6 +7,7 @@ import Mathlib.Dynamics.FixedPoints.Basic import Mathlib.GroupTheory.Perm.Option import Mathlib.Logic.Equiv.Defs import Mathlib.Logic.Equiv.Option +import Mathlib.Tactic.ApplyFun /-! # Derangements on types diff --git a/Mathlib/GroupTheory/Coprod/Basic.lean b/Mathlib/GroupTheory/Coprod/Basic.lean index 084c569e85fa9..93e15e7696015 100644 --- a/Mathlib/GroupTheory/Coprod/Basic.lean +++ b/Mathlib/GroupTheory/Coprod/Basic.lean @@ -3,7 +3,7 @@ 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.Algebra.Group.Subgroup.Basic +import Mathlib.Algebra.Group.Subgroup.Ker import Mathlib.Algebra.Group.Submonoid.Membership import Mathlib.Algebra.PUnitInstances.Algebra import Mathlib.GroupTheory.Congruence.Basic diff --git a/Mathlib/GroupTheory/FreeGroup/Basic.lean b/Mathlib/GroupTheory/FreeGroup/Basic.lean index 35f5a97da6e4a..d365be175ae24 100644 --- a/Mathlib/GroupTheory/FreeGroup/Basic.lean +++ b/Mathlib/GroupTheory/FreeGroup/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ -import Mathlib.Algebra.Group.Subgroup.Basic +import Mathlib.Algebra.Group.Subgroup.Ker import Mathlib.Data.Fintype.Basic import Mathlib.Data.List.Sublists diff --git a/Mathlib/GroupTheory/GroupAction/Basic.lean b/Mathlib/GroupTheory/GroupAction/Basic.lean index ed5ab1ed24c4a..ae7eaef28b78d 100644 --- a/Mathlib/GroupTheory/GroupAction/Basic.lean +++ b/Mathlib/GroupTheory/GroupAction/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import Mathlib.Algebra.Group.Subgroup.Basic +import Mathlib.Algebra.Group.Subgroup.Map import Mathlib.Data.Finite.Sigma import Mathlib.Data.Set.Finite.Basic import Mathlib.Data.Set.Finite.Range diff --git a/Mathlib/GroupTheory/GroupAction/FixingSubgroup.lean b/Mathlib/GroupTheory/GroupAction/FixingSubgroup.lean index f968df45e92a9..38834797ad222 100644 --- a/Mathlib/GroupTheory/GroupAction/FixingSubgroup.lean +++ b/Mathlib/GroupTheory/GroupAction/FixingSubgroup.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Antoine Chambert-Loir. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Antoine Chambert-Loir -/ -import Mathlib.Algebra.Group.Subgroup.Basic +import Mathlib.Algebra.Group.Subgroup.Lattice import Mathlib.GroupTheory.GroupAction.FixedPoints /-! diff --git a/Mathlib/GroupTheory/Perm/ClosureSwap.lean b/Mathlib/GroupTheory/Perm/ClosureSwap.lean index b00b89f6c6ce7..2256c8a2e0210 100644 --- a/Mathlib/GroupTheory/Perm/ClosureSwap.lean +++ b/Mathlib/GroupTheory/Perm/ClosureSwap.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Thomas Browning, Junyan Xu. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning, Junyan Xu -/ +import Mathlib.Algebra.Group.Subgroup.Ker import Mathlib.GroupTheory.GroupAction.Basic import Mathlib.GroupTheory.GroupAction.FixedPoints import Mathlib.GroupTheory.Perm.Support diff --git a/Mathlib/GroupTheory/Perm/Sign.lean b/Mathlib/GroupTheory/Perm/Sign.lean index d09dd86518f64..bfb55453b7c78 100644 --- a/Mathlib/GroupTheory/Perm/Sign.lean +++ b/Mathlib/GroupTheory/Perm/Sign.lean @@ -3,14 +3,15 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import Mathlib.Algebra.Group.Subgroup.Basic +import Mathlib.Algebra.Group.Conj +import Mathlib.Algebra.Group.Subgroup.Lattice import Mathlib.Algebra.Group.Submonoid.Membership import Mathlib.Algebra.Ring.Int.Units import Mathlib.Data.Finset.Fin import Mathlib.Data.Finset.Sort +import Mathlib.Data.Fintype.Prod import Mathlib.Data.Fintype.Sum import Mathlib.Data.Int.Order.Units -import Mathlib.Data.Fintype.Prod import Mathlib.GroupTheory.Perm.Support import Mathlib.Logic.Equiv.Fin import Mathlib.Tactic.NormNum.Ineq diff --git a/Mathlib/GroupTheory/PresentedGroup.lean b/Mathlib/GroupTheory/PresentedGroup.lean index ddfb60a247e99..dccb34fb884fd 100644 --- a/Mathlib/GroupTheory/PresentedGroup.lean +++ b/Mathlib/GroupTheory/PresentedGroup.lean @@ -3,6 +3,7 @@ Copyright (c) 2019 Michael Howes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Howes, Newell Jensen -/ +import Mathlib.Algebra.Group.Subgroup.Basic import Mathlib.GroupTheory.FreeGroup.Basic import Mathlib.GroupTheory.QuotientGroup.Defs diff --git a/Mathlib/GroupTheory/QuotientGroup/Defs.lean b/Mathlib/GroupTheory/QuotientGroup/Defs.lean index f193d7652942f..d33a94cc017c6 100644 --- a/Mathlib/GroupTheory/QuotientGroup/Defs.lean +++ b/Mathlib/GroupTheory/QuotientGroup/Defs.lean @@ -5,7 +5,7 @@ Authors: Kevin Buzzard, Patrick Massot -/ -- This file is to a certain extent based on `quotient_module.lean` by Johannes Hölzl. -import Mathlib.Algebra.Group.Subgroup.Basic +import Mathlib.Algebra.Group.Subgroup.Ker import Mathlib.GroupTheory.Congruence.Hom import Mathlib.GroupTheory.Coset.Defs diff --git a/Mathlib/GroupTheory/SemidirectProduct.lean b/Mathlib/GroupTheory/SemidirectProduct.lean index f3ac71a7ce4c7..f89efaef708f7 100644 --- a/Mathlib/GroupTheory/SemidirectProduct.lean +++ b/Mathlib/GroupTheory/SemidirectProduct.lean @@ -3,7 +3,8 @@ Copyright (c) 2020 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import Mathlib.Algebra.Group.Subgroup.Basic +import Mathlib.Algebra.Group.Aut +import Mathlib.Algebra.Group.Subgroup.Ker /-! # Semidirect product diff --git a/Mathlib/GroupTheory/Subgroup/Saturated.lean b/Mathlib/GroupTheory/Subgroup/Saturated.lean index 6cbc7b2c45d71..e68e30ed0f303 100644 --- a/Mathlib/GroupTheory/Subgroup/Saturated.lean +++ b/Mathlib/GroupTheory/Subgroup/Saturated.lean @@ -3,7 +3,7 @@ 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.Algebra.Group.Subgroup.Basic +import Mathlib.Algebra.Group.Subgroup.Ker import Mathlib.Algebra.NoZeroSMulDivisors.Defs /-! diff --git a/Mathlib/RingTheory/TwoSidedIdeal/Operations.lean b/Mathlib/RingTheory/TwoSidedIdeal/Operations.lean index 669cd38417774..a941b75c4e02a 100644 --- a/Mathlib/RingTheory/TwoSidedIdeal/Operations.lean +++ b/Mathlib/RingTheory/TwoSidedIdeal/Operations.lean @@ -3,12 +3,12 @@ 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.Algebra.Group.Subgroup.Map +import Mathlib.Algebra.Module.Opposite import Mathlib.Algebra.Module.Submodule.Lattice import Mathlib.RingTheory.Congruence.Opposite import Mathlib.RingTheory.Ideal.Defs import Mathlib.RingTheory.TwoSidedIdeal.Lattice -import Mathlib.Algebra.Module.Opposite /-! # Operations on two-sided ideals From c2fcf775f950f80d141260cc057afd4ffd81ae72 Mon Sep 17 00:00:00 2001 From: Jon Eugster Date: Thu, 21 Nov 2024 10:01:54 +0000 Subject: [PATCH 177/829] chore: rename `Nat.prime_def_lt''` to `Nat.prime_def` (#19255) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Rename the statement `Prime p ↔ 2 ≤ p ∧ ∀ m, m ∣ p → m = 1 ∨ m = p` to `Nat.prime_def`. Co-authored-by: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> --- Mathlib/Algebra/CharP/Defs.lean | 2 +- Mathlib/Data/Nat/Prime/Defs.lean | 11 ++++++++--- Mathlib/GroupTheory/SpecificGroups/Cyclic.lean | 2 +- Mathlib/NumberTheory/Divisors.lean | 2 +- Mathlib/NumberTheory/Fermat.lean | 2 +- scripts/no_lints_prime_decls.txt | 1 - 6 files changed, 12 insertions(+), 8 deletions(-) diff --git a/Mathlib/Algebra/CharP/Defs.lean b/Mathlib/Algebra/CharP/Defs.lean index eb523f7e90fac..093c600a4a310 100644 --- a/Mathlib/Algebra/CharP/Defs.lean +++ b/Mathlib/Algebra/CharP/Defs.lean @@ -232,7 +232,7 @@ section NoZeroDivisors variable [NoZeroDivisors R] lemma char_is_prime_of_two_le (p : ℕ) [CharP R p] (hp : 2 ≤ p) : Nat.Prime p := - suffices ∀ (d) (_ : d ∣ p), d = 1 ∨ d = p from Nat.prime_def_lt''.mpr ⟨hp, this⟩ + suffices ∀ (d) (_ : d ∣ p), d = 1 ∨ d = p from Nat.prime_def.mpr ⟨hp, this⟩ fun (d : ℕ) (hdvd : ∃ e, p = d * e) => let ⟨e, hmul⟩ := hdvd have : (p : R) = 0 := (cast_eq_zero_iff R p p).mpr (dvd_refl p) diff --git a/Mathlib/Data/Nat/Prime/Defs.lean b/Mathlib/Data/Nat/Prime/Defs.lean index e6a3e5bd1a467..30b5a96a46888 100644 --- a/Mathlib/Data/Nat/Prime/Defs.lean +++ b/Mathlib/Data/Nat/Prime/Defs.lean @@ -32,7 +32,8 @@ namespace Nat variable {n : ℕ} /-- `Nat.Prime p` means that `p` is a prime number, that is, a natural number - at least 2 whose only divisors are `p` and `1`. -/ + at least 2 whose only divisors are `p` and `1`. + The theorem `Nat.prime_def` witnesses this description of a prime number. -/ @[pp_nodot] def Prime (p : ℕ) := Irreducible p @@ -77,7 +78,8 @@ theorem Prime.eq_one_or_self_of_dvd {p : ℕ} (pp : p.Prime) (m : ℕ) (hm : m rintro rfl rw [hn, mul_one] -theorem prime_def_lt'' {p : ℕ} : Prime p ↔ 2 ≤ p ∧ ∀ m, m ∣ p → m = 1 ∨ m = p := by +@[inherit_doc Nat.Prime] +theorem prime_def {p : ℕ} : Prime p ↔ 2 ≤ p ∧ ∀ m, m ∣ p → m = 1 ∨ m = p := by refine ⟨fun h => ⟨h.two_le, h.eq_one_or_self_of_dvd⟩, fun h => ?_⟩ have h1 := Nat.one_lt_two.trans_le h.1 refine ⟨mt Nat.isUnit_iff.mp h1.ne', fun a b hab => ?_⟩ @@ -88,8 +90,11 @@ theorem prime_def_lt'' {p : ℕ} : Prime p ↔ 2 ≤ p ∧ ∀ m, m ∣ p → m · rw [hab] exact dvd_mul_right _ _ +@[deprecated (since := "2024-11-19")] +alias prime_def_lt'' := prime_def + theorem prime_def_lt {p : ℕ} : Prime p ↔ 2 ≤ p ∧ ∀ m < p, m ∣ p → m = 1 := - prime_def_lt''.trans <| + prime_def.trans <| and_congr_right fun p2 => forall_congr' fun _ => ⟨fun h l d => (h d).resolve_right (ne_of_lt l), fun h d => diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index ca2a56cc761c7..11e9dc5991b44 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -551,7 +551,7 @@ instance (priority := 100) isCyclic : IsCyclic α := by theorem prime_card [Finite α] : (Nat.card α).Prime := by have h0 : 0 < Nat.card α := Nat.card_pos obtain ⟨g, hg⟩ := IsCyclic.exists_generator (α := α) - rw [Nat.prime_def_lt''] + rw [Nat.prime_def] refine ⟨Finite.one_lt_card_iff_nontrivial.2 inferInstance, fun n hn => ?_⟩ refine (IsSimpleOrder.eq_bot_or_eq_top (Subgroup.zpowers (g ^ n))).symm.imp ?_ ?_ · intro h diff --git a/Mathlib/NumberTheory/Divisors.lean b/Mathlib/NumberTheory/Divisors.lean index e2bfeafb8e217..5d34808cc4719 100644 --- a/Mathlib/NumberTheory/Divisors.lean +++ b/Mathlib/NumberTheory/Divisors.lean @@ -391,7 +391,7 @@ theorem Prime.prod_divisors {α : Type*} [CommMonoid α] {p : ℕ} {f : ℕ → theorem properDivisors_eq_singleton_one_iff_prime : n.properDivisors = {1} ↔ n.Prime := by refine ⟨?_, ?_⟩ · intro h - refine Nat.prime_def_lt''.mpr ⟨?_, fun m hdvd => ?_⟩ + refine Nat.prime_def.mpr ⟨?_, fun m hdvd => ?_⟩ · match n with | 0 => contradiction | 1 => contradiction diff --git a/Mathlib/NumberTheory/Fermat.lean b/Mathlib/NumberTheory/Fermat.lean index 18c9ed47758f2..bdb8dd1117803 100644 --- a/Mathlib/NumberTheory/Fermat.lean +++ b/Mathlib/NumberTheory/Fermat.lean @@ -196,7 +196,7 @@ theorem prime_of_pow_sub_one_prime {a n : ℕ} (hn1 : n ≠ 1) (hP : (a ^ n - 1) rw [one_pow, hP.dvd_iff_eq (mt (Nat.sub_eq_iff_eq_add ha1.le).mp hn1), eq_comm] at h exact (pow_eq_self_iff ha1).mp (Nat.sub_one_cancel ha0 (pow_pos ha0 n) h).symm subst ha2 - refine ⟨rfl, Nat.prime_def_lt''.mpr ⟨(two_le_iff n).mpr ⟨hn0, hn1⟩, fun d hdn ↦ ?_⟩⟩ + refine ⟨rfl, Nat.prime_def.mpr ⟨(two_le_iff n).mpr ⟨hn0, hn1⟩, fun d hdn ↦ ?_⟩⟩ have hinj : ∀ x y, 2 ^ x - 1 = 2 ^ y - 1 → x = y := fun x y h ↦ Nat.pow_right_injective le_rfl (sub_one_cancel (pow_pos ha0 x) (pow_pos ha0 y) h) let h := nat_sub_dvd_pow_sub_pow (2 ^ d) 1 (n / d) diff --git a/scripts/no_lints_prime_decls.txt b/scripts/no_lints_prime_decls.txt index e04b49b8dc0f0..fd905374bfebb 100644 --- a/scripts/no_lints_prime_decls.txt +++ b/scripts/no_lints_prime_decls.txt @@ -3315,7 +3315,6 @@ Nat.Partrec.rfind' Nat.pow_lt_ascFactorial' Nat.pow_sub_lt_descFactorial' Nat.prime_def_lt' -Nat.prime_def_lt'' Nat.Prime.eq_two_or_odd' Nat.primeFactorsList_chain' Nat.Prime.not_prime_pow' From 840e02ce2768e06de7ced0a624444746590e9d99 Mon Sep 17 00:00:00 2001 From: Pietro Monticone <38562595+pitmonticone@users.noreply.github.com> Date: Thu, 21 Nov 2024 10:33:44 +0000 Subject: [PATCH 178/829] feat(Data/Int/Interval): add `Icc_eq_pair` (#19308) Upstreamed from the [Carleson](https://github.com/fpvandoorn/carleson) project. Co-authored-by: James Sundstrom --- Mathlib/Data/Int/Interval.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Mathlib/Data/Int/Interval.lean b/Mathlib/Data/Int/Interval.lean index a93f1ff49aa84..cd94ea340d1c5 100644 --- a/Mathlib/Data/Int/Interval.lean +++ b/Mathlib/Data/Int/Interval.lean @@ -129,6 +129,11 @@ theorem card_Ioc_of_le (h : a ≤ b) : (#(Ioc a b) : ℤ) = b - a := by theorem card_Ioo_of_lt (h : a < b) : (#(Ioo a b) : ℤ) = b - a - 1 := by rw [card_Ioo, sub_sub, toNat_sub_of_le h] +theorem Icc_eq_pair : Finset.Icc a (a + 1) = {a, a + 1} := by + ext + simp + omega + -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11119): removed `simp` attribute because `simpNF` says it can prove it theorem card_fintype_Icc : Fintype.card (Set.Icc a b) = (b + 1 - a).toNat := by rw [← card_Icc, Fintype.card_ofFinset] From bded1a9eb3fc57b04f6ff80b7449c55ed2a1baa4 Mon Sep 17 00:00:00 2001 From: Antoine Chambert-Loir Date: Thu, 21 Nov 2024 12:36:28 +0000 Subject: [PATCH 179/829] feat(RingTheory/DividedPowers/Basic): definition of divided powers (#15657) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define the notion of a divided power structure on an ideal of a ring. Prove basic lemmas: * the variant of the multinomial theorem in this context * a formula for products of divided powers Co-authored with: [María Inés de Frutos-Fernández](https://github.com/mariainesdff) Co-authored-by: leanprover-community-mathlib4-bot --- Mathlib.lean | 1 + Mathlib/RingTheory/DividedPowers/Basic.lean | 418 ++++++++++++++++++++ Mathlib/RingTheory/Ideal/Maps.lean | 11 + docs/references.bib | 47 +++ 4 files changed, 477 insertions(+) create mode 100644 Mathlib/RingTheory/DividedPowers/Basic.lean diff --git a/Mathlib.lean b/Mathlib.lean index 139b1e81351a5..e6b62736993a2 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4197,6 +4197,7 @@ import Mathlib.RingTheory.Derivation.ToSquareZero import Mathlib.RingTheory.DiscreteValuationRing.Basic import Mathlib.RingTheory.DiscreteValuationRing.TFAE import Mathlib.RingTheory.Discriminant +import Mathlib.RingTheory.DividedPowers.Basic import Mathlib.RingTheory.DualNumber import Mathlib.RingTheory.EisensteinCriterion import Mathlib.RingTheory.EssentialFiniteness diff --git a/Mathlib/RingTheory/DividedPowers/Basic.lean b/Mathlib/RingTheory/DividedPowers/Basic.lean new file mode 100644 index 0000000000000..b07b6b0ac8aea --- /dev/null +++ b/Mathlib/RingTheory/DividedPowers/Basic.lean @@ -0,0 +1,418 @@ +/- +Copyright (c) 2024 Antoine Chambert-Loir & María-Inés de Frutos—Fernández. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Antoine Chambert-Loir, María-Inés de Frutos—Fernández +-/ + +import Mathlib.RingTheory.PowerSeries.Basic +import Mathlib.Combinatorics.Enumerative.Bell +import Mathlib.Data.Nat.Choose.Multinomial + +/-! # Divided powers + +Let `A` be a commutative (semi)ring and `I` be an ideal of `A`. +A *divided power* structure on `I` is the datum of operations `a n ↦ dpow a n` +satisfying relations that model the intuitive formula `dpow n a = a ^ n / n !` and +collected by the structure `DividedPowers`. The list of axioms is embedded in the structure: +To avoid coercions, we rather consider `DividedPowers.dpow : ℕ → A → A`, extended by 0. + +* `DividedPowers.dpow_null` asserts that `dpow n x = 0` for `x ∉ I` + +* `DividedPowers.dpow_mem` : `dpow n x ∈ I` for `n ≠ 0` +For `x y : A` and `m n : ℕ` such that `x ∈ I` and `y ∈ I`, one has +* `DividedPowers.dpow_zero` : `dpow 0 x = 1` +* `DividedPowers.dpow_one` : `dpow 1 x = 1` +* `DividedPowers.dpow_add` : `dpow n (x + y) = +(antidiagonal n).sum fun k ↦ dpow k.1 x * dpow k.2 y`, +this is the binomial theorem without binomial coefficients. +* `DividedPowers.dpow_mul`: `dpow n (a * x) = a ^ n * dpow n x` +* `DividedPowers.mul_dpow` : `dpow m x * dpow n x = choose (m + n) m * dpow (m + n) x` +* `DividedPowers.dpow_comp` : `dpow m (dpow n x) = uniformBell m n * dpow (m * n) x` + +* `DividedPowers.dividedPowersBot` : the trivial divided powers structure on the zero ideal + +* `DividedPowers.prod_dpow`: a product of divided powers is a multinomial coefficients +times a divided power + +* `DividedPowers.dpow_sum`: the multinomial theorem for divided powers, +without multinomial coefficients. + +* `DividedPowers.ofRingEquiv`: transfer divided powers along `RingEquiv` + +* `DividedPowers.equiv`: the equivalence `DividedPowers I ≃ DividedPowers J`, +for `e : R ≃+* S`, and `I : Ideal R`, `J : Ideal S` such that `I.map e = J` + +* `DividedPowers.exp`: the power series `Σ (dpow n a) X ^n` + +* `DividedPowers.exp_add`: its multiplicativity + +## References + +* [P. Berthelot (1974), *Cohomologie cristalline des schémas de +caractéristique $p$ > 0*][Berthelot-1974] + +* [P. Berthelot and A. Ogus (1978), *Notes on crystalline +cohomology*][BerthelotOgus-1978] + +* [N. Roby (1963), *Lois polynomes et lois formelles en théorie des +modules*][Roby-1963] + +* [N. Roby (1965), *Les algèbres à puissances dividées*][Roby-1965] + +## Discussion + +* In practice, one often has a single such structure to handle on a given ideal, +but several ideals of the same ring might be considered. +Without any explicit mention of the ideal, it is not clear whether such structures +should be provided as instances. + +* We do not provide any notation such as `a ^[n]` for `dpow a n`. + +-/ + +open Finset Nat Ideal + +section DividedPowersDefinition +/- ## Definition of divided powers -/ + +variable {A : Type*} [CommSemiring A] (I : Ideal A) + +/-- The divided power structure on an ideal I of a commutative ring A -/ +structure DividedPowers where + /-- The divided power function underlying a divided power structure -/ + dpow : ℕ → A → A + dpow_null : ∀ {n x} (_ : x ∉ I), dpow n x = 0 + dpow_zero : ∀ {x} (_ : x ∈ I), dpow 0 x = 1 + dpow_one : ∀ {x} (_ : x ∈ I), dpow 1 x = x + dpow_mem : ∀ {n x} (_ : n ≠ 0) (_ : x ∈ I), dpow n x ∈ I + dpow_add : ∀ (n) {x y} (_ : x ∈ I) (_ : y ∈ I), + dpow n (x + y) = (antidiagonal n).sum fun k ↦ dpow k.1 x * dpow k.2 y + dpow_mul : ∀ (n) {a : A} {x} (_ : x ∈ I), + dpow n (a * x) = a ^ n * dpow n x + mul_dpow : ∀ (m n) {x} (_ : x ∈ I), + dpow m x * dpow n x = choose (m + n) m * dpow (m + n) x + dpow_comp : ∀ (m) {n x} (_ : n ≠ 0) (_ : x ∈ I), + dpow m (dpow n x) = uniformBell m n * dpow (m * n) x + +variable (A) in +/-- The canonical `DividedPowers` structure on the zero ideal -/ +def dividedPowersBot [DecidableEq A] : DividedPowers (⊥ : Ideal A) where + dpow n a := ite (a = 0 ∧ n = 0) 1 0 + dpow_null {n a} ha := by + simp only [mem_bot] at ha + dsimp + rw [if_neg] + exact not_and_of_not_left (n = 0) ha + dpow_zero {a} ha := by + rw [mem_bot.mp ha] + simp only [and_self, ite_true] + dpow_one {a} ha := by + simp [mem_bot.mp ha] + dpow_mem {n a} hn _ := by + simp only [mem_bot, ite_eq_right_iff, and_imp] + exact fun _ a ↦ False.elim (hn a) + dpow_add n a b ha hb := by + rw [mem_bot.mp ha, mem_bot.mp hb, add_zero] + simp only [true_and, ge_iff_le, tsub_eq_zero_iff_le, mul_ite, mul_one, mul_zero] + split_ifs with h + · simp [h] + · symm + apply sum_eq_zero + intro i hi + simp only [mem_antidiagonal] at hi + split_ifs with h2 h1 + · rw [h1, h2, add_zero] at hi + exfalso + exact h hi.symm + · rfl + · rfl + dpow_mul n {a x} hx := by + rw [mem_bot.mp hx] + simp only [mul_zero, true_and, mul_ite, mul_one] + by_cases hn : n = 0 + · rw [if_pos hn, hn, if_pos rfl, _root_.pow_zero] + · simp only [if_neg hn] + mul_dpow m n {x} hx := by + rw [mem_bot.mp hx] + simp only [true_and, mul_ite, mul_one, mul_zero, add_eq_zero] + by_cases hn : n = 0 + · simp only [hn, ite_true, and_true, add_zero, choose_self, cast_one] + · rw [if_neg hn, if_neg] + exact not_and_of_not_right (m = 0) hn + dpow_comp m {n a} hn ha := by + rw [mem_bot.mp ha] + simp only [true_and, ite_eq_right_iff, _root_.mul_eq_zero, mul_ite, mul_one, mul_zero] + by_cases hm: m = 0 + · simp only [hm, and_true, true_or, ite_true, uniformBell_zero_left, cast_one] + rw [if_pos] + exact fun h ↦ False.elim (hn h) + · simp only [hm, and_false, ite_false, false_or, if_neg hn] + +instance [DecidableEq A] : Inhabited (DividedPowers (⊥ : Ideal A)) := + ⟨dividedPowersBot A⟩ + +/-- The coercion from the divided powers structures to functions -/ +instance : CoeFun (DividedPowers I) fun _ ↦ ℕ → A → A := ⟨fun hI ↦ hI.dpow⟩ + +variable {I} in +@[ext] +theorem DividedPowers.ext (hI : DividedPowers I) (hI' : DividedPowers I) + (h_eq : ∀ (n : ℕ) {x : A} (_ : x ∈ I), hI.dpow n x = hI'.dpow n x) : + hI = hI' := by + obtain ⟨hI, h₀, _⟩ := hI + obtain ⟨hI', h₀', _⟩ := hI' + simp only [mk.injEq] + ext n x + by_cases hx : x ∈ I + · exact h_eq n hx + · rw [h₀ hx, h₀' hx] + +theorem DividedPowers.coe_injective : + Function.Injective (fun (h : DividedPowers I) ↦ (h : ℕ → A → A)) := fun hI hI' h ↦ by + ext n x + exact congr_fun (congr_fun h n) x + +end DividedPowersDefinition + +namespace DividedPowers + +section BasicLemmas + +/- ## Basic lemmas for divided powers -/ + +variable {A : Type*} [CommSemiring A] {I : Ideal A} {a b : A} + +/-- Variant of `DividedPowers.dpow_add` with a sum on `range (n + 1)` -/ +theorem dpow_add' (hI : DividedPowers I) (n : ℕ) (ha : a ∈ I) (hb : b ∈ I) : + hI.dpow n (a + b) = (range (n + 1)).sum fun k ↦ hI.dpow k a * hI.dpow (n - k) b := by + rw [hI.dpow_add n ha hb, sum_antidiagonal_eq_sum_range_succ_mk] + +/-- The exponential series of an element in the context of divided powers, +`Σ (dpow n a) X ^ n` -/ +def exp (hI : DividedPowers I) (a : A) : PowerSeries A := + PowerSeries.mk fun n ↦ hI.dpow n a + +/-- A more general of `DividedPowers.exp_add` -/ +theorem exp_add' (dp : ℕ → A → A) + (dp_add : ∀ n, dp n (a + b) = (antidiagonal n).sum fun k ↦ dp k.1 a * dp k.2 b) : + PowerSeries.mk (fun n ↦ dp n (a + b)) = + (PowerSeries.mk fun n ↦ dp n a) * (PowerSeries.mk fun n ↦ dp n b) := by + ext n + simp only [exp, PowerSeries.coeff_mk, PowerSeries.coeff_mul, dp_add n, + sum_antidiagonal_eq_sum_range_succ_mk] + +theorem exp_add (hI : DividedPowers I) (ha : a ∈ I) (hb : b ∈ I) : + hI.exp (a + b) = hI.exp a * hI.exp b := + exp_add' _ (fun n ↦ hI.dpow_add n ha hb) + +variable (hI : DividedPowers I) + +/- ## Rewriting lemmas -/ + +theorem dpow_smul (n : ℕ) (ha : a ∈ I) : + hI.dpow n (b • a) = b ^ n • hI.dpow n a := by + simp only [smul_eq_mul, hI.dpow_mul, ha] + +theorem dpow_mul_right (n : ℕ) (ha : a ∈ I) : + hI.dpow n (a * b) = hI.dpow n a * b ^ n := by + rw [mul_comm, hI.dpow_mul n ha, mul_comm] + +theorem dpow_smul_right (n : ℕ) (ha : a ∈ I) : + hI.dpow n (a • b) = hI.dpow n a • b ^ n := by + rw [smul_eq_mul, hI.dpow_mul_right n ha, smul_eq_mul] + +theorem factorial_mul_dpow_eq_pow (n : ℕ) (ha : a ∈ I) : + (n ! : A) * hI.dpow n a = a ^ n := by + induction n with + | zero => rw [factorial_zero, cast_one, one_mul, pow_zero, hI.dpow_zero ha] + | succ n ih => + rw [factorial_succ, mul_comm (n + 1)] + nth_rewrite 1 [← (n + 1).choose_one_right] + rw [← choose_symm_add, cast_mul, mul_assoc, + ← hI.mul_dpow n 1 ha, ← mul_assoc, ih, hI.dpow_one ha, pow_succ, mul_comm] + +theorem dpow_eval_zero {n : ℕ} (hn : n ≠ 0) : hI.dpow n 0 = 0 := by + rw [← MulZeroClass.mul_zero (0 : A), hI.dpow_mul n I.zero_mem, + zero_pow hn, zero_mul, zero_mul] + +/-- If an element of a divided power ideal is killed by multiplication +by some nonzero integer `n`, then its `n`th power is zero. + +Proposition 1.2.7 of [Berthelot-1974], part (i). -/ +theorem nilpotent_of_mem_dpIdeal {n : ℕ} (hn : n ≠ 0) (hnI : ∀ {y} (_ : y ∈ I), n • y = 0) + (hI : DividedPowers I) (ha : a ∈ I) : a ^ n = 0 := by + have h_fac : (n ! : A) * hI.dpow n a = + n • ((n - 1)! : A) * hI.dpow n a := by + rw [nsmul_eq_mul, ← cast_mul, mul_factorial_pred (Nat.pos_of_ne_zero hn)] + rw [← hI.factorial_mul_dpow_eq_pow n ha, h_fac, smul_mul_assoc] + exact hnI (I.mul_mem_left ((n - 1)! : A) (hI.dpow_mem hn ha)) + +/-- If J is another ideal of A with divided powers, +then the divided powers of I and J coincide on I • J + +[Berthelot-1974], 1.6.1 (ii) -/ +theorem coincide_on_smul {J : Ideal A} (hJ : DividedPowers J) {n : ℕ} (ha : a ∈ I • J) : + hI.dpow n a = hJ.dpow n a := by + induction ha using Submodule.smul_induction_on' generalizing n with + | smul a ha b hb => + rw [Algebra.id.smul_eq_mul, hJ.dpow_mul n hb, mul_comm a b, hI.dpow_mul n ha, ← + hJ.factorial_mul_dpow_eq_pow n hb, ← hI.factorial_mul_dpow_eq_pow n ha] + ring + | add x hx y hy hx' hy' => + rw [hI.dpow_add n (mul_le_right hx) (mul_le_right hy), + hJ.dpow_add n (mul_le_left hx) (mul_le_left hy)] + apply sum_congr rfl + intro k _ + rw [hx', hy'] + +/-- A product of divided powers is a multinomial coefficient times the divided power + +[Roby-1965], formula (III') -/ +theorem prod_dpow {ι : Type*} {s : Finset ι} (n : ι → ℕ) (ha : a ∈ I) : + (s.prod fun i ↦ hI.dpow (n i) a) = multinomial s n * hI.dpow (s.sum n) a := by + classical + induction s using Finset.induction with + | empty => + simp only [prod_empty, multinomial_empty, cast_one, sum_empty, one_mul] + rw [hI.dpow_zero ha] + | insert hi hrec => + rw [prod_insert hi, hrec, ← mul_assoc, mul_comm (hI.dpow (n _) a), + mul_assoc, mul_dpow _ _ _ ha, ← sum_insert hi, ← mul_assoc] + apply congr_arg₂ _ _ rfl + rw [multinomial_insert hi, mul_comm, cast_mul, sum_insert hi] + +-- TODO : can probably be simplified using `DividedPowers.exp` + +/-- Lemma towards `dpow_sum` when we only have partial information on a divided power ideal -/ +theorem dpow_sum' {M : Type*} [AddCommMonoid M] {I : AddSubmonoid M} (dpow : ℕ → M → A) + (dpow_zero : ∀ {x} (_ : x ∈ I), dpow 0 x = 1) + (dpow_add : ∀ (n) {x y} (_ : x ∈ I) (_ : y ∈ I), + dpow n (x + y) = (antidiagonal n).sum fun k ↦ dpow k.1 x * dpow k.2 y) + (dpow_eval_zero : ∀ {n : ℕ} (_ : n ≠ 0), dpow n 0 = 0) + {ι : Type*} [DecidableEq ι] {s : Finset ι} {x : ι → M} (hx : ∀ i ∈ s, x i ∈ I) (n : ℕ) : + dpow n (s.sum x) = (s.sym n).sum fun k ↦ s.prod fun i ↦ dpow (Multiset.count i k) (x i) := by + simp only [sum_antidiagonal_eq_sum_range_succ_mk] at dpow_add + induction s using Finset.induction generalizing n with + | empty => + simp only [sum_empty, prod_empty, sum_const, nsmul_eq_mul, mul_one] + by_cases hn : n = 0 + · rw [hn] + rw [dpow_zero I.zero_mem] + simp only [sym_zero, card_singleton, cast_one] + · rw [dpow_eval_zero hn, eq_comm, ← cast_zero] + apply congr_arg + rw [card_eq_zero, sym_eq_empty] + exact ⟨hn, rfl⟩ + | @insert a s ha ih => + -- This should be golfable using `Finset.symInsertEquiv` + have hx' : ∀ i, i ∈ s → x i ∈ I := fun i hi ↦ hx i (mem_insert_of_mem hi) + simp_rw [sum_insert ha, + dpow_add n (hx a (mem_insert_self a s)) (I.sum_mem fun i ↦ hx' i), + sum_range, ih hx', mul_sum, sum_sigma', eq_comm] + apply sum_bij' + (fun m _ ↦ m.filterNe a) + (fun m _ ↦ m.2.fill a m.1) + (fun m hm ↦ mem_sigma.2 ⟨mem_univ _, _⟩) + (fun m hm ↦ by + simp only [succ_eq_add_one, mem_sym_iff, mem_insert, Sym.mem_fill_iff] + simp only [mem_sigma, mem_univ, mem_sym_iff, true_and] at hm + intro b + apply Or.imp (fun h ↦ h.2) (fun h ↦ hm b h)) + (fun m _ ↦ m.fill_filterNe a) + · intro m hm + simp only [mem_sigma, mem_univ, mem_sym_iff, true_and] at hm + exact Sym.filter_ne_fill a m fun a_1 ↦ ha (hm a a_1) + · intro m hm + simp only [mem_sym_iff, mem_insert] at hm + rw [prod_insert ha] + apply congr_arg₂ _ rfl + apply prod_congr rfl + intro i hi + apply congr_arg₂ _ _ rfl + conv_lhs => rw [← m.fill_filterNe a] + exact Sym.count_coe_fill_of_ne (ne_of_mem_of_not_mem hi ha) + · intro m hm + convert sym_filterNe_mem a hm + rw [erase_insert ha] + +/-- A “multinomial” theorem for divided powers — without multinomial coefficients -/ +theorem dpow_sum {ι : Type*} [DecidableEq ι] {s : Finset ι} {x : ι → A} + (hx : ∀ i ∈ s, x i ∈ I) (n : ℕ) : + hI.dpow n (s.sum x) = + (s.sym n).sum fun k ↦ s.prod fun i ↦ hI.dpow (Multiset.count i k) (x i) := + dpow_sum' hI.dpow hI.dpow_zero hI.dpow_add hI.dpow_eval_zero hx n + +end BasicLemmas + +section Equiv +/- ## Relation of divided powers with ring equivalences -/ + +variable {A B : Type*} [CommSemiring A] {I : Ideal A} [CommSemiring B] {J : Ideal B} + {e : A ≃+* B} (h : I.map e = J) + +/-- Transfer divided powers under an equivalence -/ +def ofRingEquiv (hI : DividedPowers I) : DividedPowers J where + dpow n b := e (hI.dpow n (e.symm b)) + dpow_null {n} {x} hx := by + rw [EmbeddingLike.map_eq_zero_iff, hI.dpow_null] + rwa [symm_apply_mem_of_equiv_iff, h] + dpow_zero {x} hx := by + rw [EmbeddingLike.map_eq_one_iff, hI.dpow_zero] + rwa [symm_apply_mem_of_equiv_iff, h] + dpow_one {x} hx := by + simp only + rw [dpow_one, RingEquiv.apply_symm_apply] + rwa [I.symm_apply_mem_of_equiv_iff, h] + dpow_mem {n} {x} hn hx := by + simp only + rw [← h, I.apply_mem_of_equiv_iff] + apply hI.dpow_mem hn + rwa [I.symm_apply_mem_of_equiv_iff, h] + dpow_add n {x y} hx hy := by + simp only [map_add] + rw [hI.dpow_add n (symm_apply_mem_of_equiv_iff.mpr (h ▸ hx)) + (symm_apply_mem_of_equiv_iff.mpr (h ▸ hy))] + simp only [map_sum, _root_.map_mul] + dpow_mul n {a x} hx := by + simp only [_root_.map_mul] + rw [hI.dpow_mul n (symm_apply_mem_of_equiv_iff.mpr (h ▸ hx))] + rw [_root_.map_mul, map_pow] + simp only [RingEquiv.apply_symm_apply] + mul_dpow m n {x} hx := by + simp only + rw [← _root_.map_mul, hI.mul_dpow, _root_.map_mul] + · simp only [map_natCast] + · rwa [symm_apply_mem_of_equiv_iff, h] + dpow_comp m {n x} hn hx := by + simp only [RingEquiv.symm_apply_apply] + rw [hI.dpow_comp _ hn] + · simp only [_root_.map_mul, map_natCast] + · rwa [symm_apply_mem_of_equiv_iff, h] + +@[simp] +theorem ofRingEquiv_dpow (hI : DividedPowers I) (n : ℕ) (b : B) : + (ofRingEquiv h hI).dpow n b = e (hI.dpow n (e.symm b)) := rfl + +theorem ofRingEquiv_dpow_apply (hI : DividedPowers I) (n : ℕ) (a : A) : + (ofRingEquiv h hI).dpow n (e a) = e (hI.dpow n a) := by + simp + +/-- Transfer divided powers under an equivalence (Equiv version) -/ +def equiv : DividedPowers I ≃ DividedPowers J where + toFun := ofRingEquiv h + invFun := ofRingEquiv (show map e.symm J = I by rw [← h]; exact I.map_of_equiv e) + left_inv := fun hI ↦ by ext n a; simp [ofRingEquiv] + right_inv := fun hJ ↦ by ext n b; simp [ofRingEquiv] + +theorem equiv_apply (hI : DividedPowers I) (n : ℕ) (b : B) : + (equiv h hI).dpow n b = e (hI.dpow n (e.symm b)) := rfl + +/-- Variant of `DividedPowers.equiv_apply` -/ +theorem equiv_apply' (hI : DividedPowers I) (n : ℕ) (a : A) : + (equiv h hI).dpow n (e a) = e (hI.dpow n a) := + ofRingEquiv_dpow_apply h hI n a + +end Equiv + +end DividedPowers + diff --git a/Mathlib/RingTheory/Ideal/Maps.lean b/Mathlib/RingTheory/Ideal/Maps.lean index 7a67ded065aff..e2a631a34c0a9 100644 --- a/Mathlib/RingTheory/Ideal/Maps.lean +++ b/Mathlib/RingTheory/Ideal/Maps.lean @@ -362,9 +362,20 @@ theorem comap_symm {I : Ideal R} (f : R ≃+* S) : I.comap f.symm = I.map f := /-- If `f : R ≃+* S` is a ring isomorphism and `I : Ideal R`, then `map f.symm I = comap f I`. -/ @[simp] + theorem map_symm {I : Ideal S} (f : R ≃+* S) : I.map f.symm = I.comap f := map_comap_of_equiv (RingEquiv.symm f) +@[simp] +theorem symm_apply_mem_of_equiv_iff {I : Ideal R} {f : R ≃+* S} {y : S} : + f.symm y ∈ I ↔ y ∈ I.map f := by + rw [← comap_symm, mem_comap] + +@[simp] +theorem apply_mem_of_equiv_iff {I : Ideal R} {f : R ≃+* S} {x : R} : + f x ∈ I.map f ↔ x ∈ I := by + rw [← comap_symm, Ideal.mem_comap, f.symm_apply_apply] + theorem mem_map_of_equiv {E : Type*} [EquivLike E R S] [RingEquivClass E R S] (e : E) {I : Ideal R} (y : S) : y ∈ map e I ↔ ∃ x ∈ I, e x = y := by constructor diff --git a/docs/references.bib b/docs/references.bib index 6603b6a5466fe..392120a19de1e 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -305,6 +305,31 @@ @Article{ bernstein1912 number = {1–2} } +@Book{ Berthelot-1974, + author = {Berthelot, Pierre}, + year = {1974}, + series = {Lecture {{Notes}} in {{Mathematics}}}, + number = {407}, + publisher = {sv}, + keywords = {nosource}, + file = {/Users/chambert/Zotero/storage/WHF5R9LQ/Berthelot - 1974 - + Cohomologie cristalline des schémas de caractérist.djvu}, + title = {Cohomologie Cristalline Des Schémas de Caractéristique + {$p>0$}} +} + +@Book{ BerthelotOgus-1978, + title = {Notes on Crystalline Cohomology}, + author = {Berthelot, Pierre and Ogus, Arthur}, + year = {1978}, + series = {Math. {{Notes}}}, + number = {21}, + publisher = {pup}, + keywords = {nosource}, + file = {/Users/chambert/Zotero/storage/I3FNJEDB/Berthelot et Ogus + - 1978 - Notes on crystalline cohomology.djvu} +} + @InProceedings{ beylin1996, author = "Beylin, Ilya and Dybjer, Peter", editor = "Berardi, Stefano and Coppo, Mario", @@ -3036,6 +3061,28 @@ @Misc{ RisingSea url = "https://math.stanford.edu/~vakil/216blog/" } +@Article{ Roby-1963, + title = {Lois polynomes et lois formelles en théorie des modules}, + author = {Roby, Norbert}, + year = {1963}, + journal = {Annales scientifiques de l'École Normale Supérieure}, + volume = {80}, + number = {3}, + pages = {213--348}, + doi = {10.24033/asens.1124}, + urldate = {2022-09-15}, + langid = {french} +} + +@Article{ Roby-1965, + title = {Les Algèbres à Puissances Divisées}, + author = {Roby, Norbert}, + year = {1965}, + journal = {Bulletin des Sciences Mathématiques. Deuxième Série}, + volume = {89}, + pages = {75--91} +} + @Article{ Rosenlicht_1972, author = {Rosenlicht, Maxwell}, title = {Integration in finite terms}, From 29b03c858f01fe6ef2627bb4fce7d71e99fcc517 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Thu, 21 Nov 2024 13:54:59 +0000 Subject: [PATCH 180/829] chore: update Mathlib dependencies 2024-11-21 (#19333) 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 8fc2bedf751bd..9d7aff8b1542f 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "a822446d61ad7e7f5e843365c7041c326553050a", + "rev": "33d7f346440869364a2cd077bde8cebf73243aaa", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From d59b1904064942e8d375bc7091a222738e2f6539 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Thu, 21 Nov 2024 14:50:51 +0000 Subject: [PATCH 181/829] feat(NumberTheory/LSeries/Convergence): add lemma on real-valued series (#19334) This is part of the series of PRs leading to Dirichlet's Theorem on primes in AP. See [here](https://leanprover.zulipchat.com/#narrow/channel/144837-PR-reviews/topic/Prerequisites.20for.20PNT.20and.20Dirichlet's.20Thm/near/483303184) on Zulip. --- Mathlib/NumberTheory/LSeries/Convergence.lean | 18 ++++++++++++++++++ 1 file changed, 18 insertions(+) diff --git a/Mathlib/NumberTheory/LSeries/Convergence.lean b/Mathlib/NumberTheory/LSeries/Convergence.lean index c5a161fa9aae9..5310dc93a15ed 100644 --- a/Mathlib/NumberTheory/LSeries/Convergence.lean +++ b/Mathlib/NumberTheory/LSeries/Convergence.lean @@ -119,3 +119,21 @@ lemma LSeries.abscissaOfAbsConv_le_one_of_isBigO_one {f : ℕ → ℂ} (h : f =O convert abscissaOfAbsConv_le_of_isBigO_rpow (x := 0) ?_ · simp only [EReal.coe_zero, zero_add] · simpa only [Real.rpow_zero] using h + +/-- If `f` is real-valued and `x` is strictly greater than the abscissa of absolute convergence +of `f`, then the real series `∑' n, f n / n ^ x` converges. -/ +lemma LSeries.summable_real_of_abscissaOfAbsConv_lt {f : ℕ → ℝ} {x : ℝ} + (h : abscissaOfAbsConv (f ·) < x) : + Summable fun n : ℕ ↦ f n / (n : ℝ) ^ x := by + have h' : abscissaOfAbsConv (f ·) < (x : ℂ).re := by simpa only [ofReal_re] using h + have := LSeriesSummable_of_abscissaOfAbsConv_lt_re h' + rw [LSeriesSummable, show term _ _ = fun n ↦ _ from rfl] at this + conv at this => + enter [1, n] + rw [term_def, show (n : ℂ) = (n : ℝ) from rfl, ← ofReal_cpow n.cast_nonneg, ← ofReal_div, + show (0 : ℂ) = (0 : ℝ) from rfl, ← apply_ite] + rw [summable_ofReal] at this + refine this.congr_cofinite ?_ + filter_upwards [Set.Finite.compl_mem_cofinite <| Set.finite_singleton 0] with n hn + simp only [Set.mem_compl_iff, Set.mem_singleton_iff] at hn + exact if_neg hn From e2d0dfcecaa519889ccdb80b7ff61469c501d795 Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Thu, 21 Nov 2024 15:34:01 +0000 Subject: [PATCH 182/829] feat: add `IsFractionRing.lift[AlgHom]_fieldRange[_eq_of_range_eq]` (#19286) Which are direct applications of `IsFractionRing.ringHom_fieldRange_eq_of_comp_eq`. Also fix some docstrings mentioning "integral domain". --- Mathlib/FieldTheory/Adjoin.lean | 15 +++++++ .../RingTheory/Localization/FractionRing.lean | 39 ++++++++++++------- 2 files changed, 41 insertions(+), 13 deletions(-) diff --git a/Mathlib/FieldTheory/Adjoin.lean b/Mathlib/FieldTheory/Adjoin.lean index e53c89cf4befb..9ccf8c539d1f9 100644 --- a/Mathlib/FieldTheory/Adjoin.lean +++ b/Mathlib/FieldTheory/Adjoin.lean @@ -1663,6 +1663,21 @@ theorem algHom_fieldRange_eq_of_comp_eq_of_range_eq refine ringHom_fieldRange_eq_of_comp_eq_of_range_eq h ?_ rw [← Algebra.adjoin_eq_ring_closure, ← hs]; rfl +variable [IsScalarTower F A K] + +/-- The image of `IsFractionRing.liftAlgHom` is the intermediate field generated by the image +of the algebra hom. -/ +theorem liftAlgHom_fieldRange (hg : Function.Injective g) : + (liftAlgHom hg : K →ₐ[F] L).fieldRange = IntermediateField.adjoin F g.range := + algHom_fieldRange_eq_of_comp_eq (by ext; simp) + +/-- The image of `IsFractionRing.liftAlgHom` is the intermediate field generated by `s`, +if the image of the algebra hom is the subalgebra generated by `s`. -/ +theorem liftAlgHom_fieldRange_eq_of_range_eq (hg : Function.Injective g) + {s : Set L} (hs : g.range = Algebra.adjoin F s) : + (liftAlgHom hg : K →ₐ[F] L).fieldRange = IntermediateField.adjoin F s := + algHom_fieldRange_eq_of_comp_eq_of_range_eq (by ext; simp) hs + end IsFractionRing set_option linter.style.longFile 1700 diff --git a/Mathlib/RingTheory/Localization/FractionRing.lean b/Mathlib/RingTheory/Localization/FractionRing.lean index fb9b40025a8df..155c5ec2966a3 100644 --- a/Mathlib/RingTheory/Localization/FractionRing.lean +++ b/Mathlib/RingTheory/Localization/FractionRing.lean @@ -39,7 +39,7 @@ variable [Algebra R S] {P : Type*} [CommRing P] variable {A : Type*} [CommRing A] (K : Type*) -- TODO: should this extend `Algebra` instead of assuming it? -/-- `IsFractionRing R K` states `K` is the field of fractions of an integral domain `R`. -/ +/-- `IsFractionRing R K` states `K` is the ring of fractions of a commutative ring `R`. -/ abbrev IsFractionRing [CommRing K] [Algebra R K] := IsLocalization (nonZeroDivisors R) K @@ -186,8 +186,8 @@ theorem mk'_eq_one_iff_eq {x : A} {y : nonZeroDivisors A} : mk' K x y = 1 ↔ x section Subfield variable (A K) in -/-- If `A` is a ring with fraction field `K`, then the subfield of `K` generated by the image of -`algebraMap A K` is equal to the whole field `K`. -/ +/-- If `A` is a commutative ring with fraction field `K`, then the subfield of `K` generated by +the image of `algebraMap A K` is equal to the whole field `K`. -/ theorem closure_range_algebraMap : Subfield.closure (Set.range (algebraMap A K)) = ⊤ := top_unique fun z _ ↦ by obtain ⟨_, _, -, rfl⟩ := div_surjective (A := A) z @@ -195,16 +195,16 @@ theorem closure_range_algebraMap : Subfield.closure (Set.range (algebraMap A K)) variable {L : Type*} [Field L] {g : A →+* L} {f : K →+* L} -/-- If `A` is a ring with fraction field `K`, `L` is a field, `g : A →+* L` lifts to `f : K →+* L`, -then the image of `f` is the field generated by the image of `g`. -/ +/-- If `A` is a commutative ring with fraction field `K`, `L` is a field, `g : A →+* L` lifts to +`f : K →+* L`, then the image of `f` is the subfield generated by the image of `g`. -/ theorem ringHom_fieldRange_eq_of_comp_eq (h : RingHom.comp f (algebraMap A K) = g) : f.fieldRange = Subfield.closure g.range := by rw [f.fieldRange_eq_map, ← closure_range_algebraMap A K, f.map_field_closure, ← Set.range_comp, ← f.coe_comp, h, g.coe_range] -/-- If `A` is a ring with fraction field `K`, `L` is a field, `g : A →+* L` lifts to `f : K →+* L`, -`s` is a set such that the image of `g` is the subring generated by `s`, -then the image of `f` is the field generated by `s`. -/ +/-- If `A` is a commutative ring with fraction field `K`, `L` is a field, `g : A →+* L` lifts to +`f : K →+* L`, `s` is a set such that the image of `g` is the subring generated by `s`, +then the image of `f` is the subfield generated by `s`. -/ theorem ringHom_fieldRange_eq_of_comp_eq_of_range_eq (h : RingHom.comp f (algebraMap A K) = g) {s : Set L} (hs : g.range = Subring.closure s) : f.fieldRange = Subfield.closure s := by rw [ringHom_fieldRange_eq_of_comp_eq h, hs] @@ -215,7 +215,7 @@ end Subfield open Function -/-- Given an integral domain `A` with field of fractions `K`, +/-- Given a commutative ring `A` with field of fractions `K`, and an injective ring hom `g : A →+* L` where `L` is a field, we get a field hom sending `z : K` to `g x * (g y)⁻¹`, where `(x, y) : A × (NonZeroDivisors A)` are such that `z = f x * (f y)⁻¹`. -/ @@ -241,7 +241,7 @@ theorem liftAlgHom_apply : liftAlgHom hg x = lift hg x := rfl end liftAlgHom -/-- Given an integral domain `A` with field of fractions `K`, +/-- Given a commutative ring `A` with field of fractions `K`, and an injective ring hom `g : A →+* L` where `L` is a field, the field hom induced from `K` to `L` maps `x` to `g x` for all `x : A`. -/ @@ -249,15 +249,28 @@ the field hom induced from `K` to `L` maps `x` to `g x` for all theorem lift_algebraMap (hg : Injective g) (x) : lift hg (algebraMap A K x) = g x := lift_eq _ _ -/-- Given an integral domain `A` with field of fractions `K`, +/-- The image of `IsFractionRing.lift` is the subfield generated by the image +of the ring hom. -/ +theorem lift_fieldRange (hg : Injective g) : + (lift hg : K →+* L).fieldRange = Subfield.closure g.range := + ringHom_fieldRange_eq_of_comp_eq (by ext; simp) + +/-- The image of `IsFractionRing.lift` is the subfield generated by `s`, if the image +of the ring hom is the subring generated by `s`. -/ +theorem lift_fieldRange_eq_of_range_eq (hg : Injective g) + {s : Set L} (hs : g.range = Subring.closure s) : + (lift hg : K →+* L).fieldRange = Subfield.closure s := + ringHom_fieldRange_eq_of_comp_eq_of_range_eq (by ext; simp) hs + +/-- Given a commutative ring `A` with field of fractions `K`, and an injective ring hom `g : A →+* L` where `L` is a field, field hom induced from `K` to `L` maps `f x / f y` to `g x / g y` for all `x : A, y ∈ NonZeroDivisors A`. -/ theorem lift_mk' (hg : Injective g) (x) (y : nonZeroDivisors A) : lift hg (mk' K x y) = g x / g y := by simp only [mk'_eq_div, map_div₀, lift_algebraMap] -/-- Given integral domains `A, B` with fields of fractions `K`, `L` -and an injective ring hom `j : A →+* B`, we get a field hom +/-- Given commutative rings `A, B` where `B` is an integral domain, with fraction rings `K`, `L` +and an injective ring hom `j : A →+* B`, we get a ring hom sending `z : K` to `g (j x) * (g (j y))⁻¹`, where `(x, y) : A × (NonZeroDivisors A)` are such that `z = f x * (f y)⁻¹`. -/ noncomputable def map {A B K L : Type*} [CommRing A] [CommRing B] [IsDomain B] [CommRing K] From d1c99f29771dcd3c24fa98f38a9baefa5598f548 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Martin=20Dvo=C5=99=C3=A1k?= Date: Thu, 21 Nov 2024 15:34:02 +0000 Subject: [PATCH 183/829] docs: `Matrix.IsTotallyUnimodular` docstring (#19338) --- .../LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean index 1c03274799f13..3c81f42c1c1b9 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean @@ -28,7 +28,8 @@ namespace Matrix variable {m n R : Type*} [CommRing R] -/-- Is the matrix `A` totally unimodular? -/ +/-- `A.IsTotallyUnimodular` means that every square submatrix of `A` (not necessarily contiguous) +has determinant `0` or `1` or `-1`. -/ def IsTotallyUnimodular (A : Matrix m n R) : Prop := ∀ k : ℕ, ∀ f : Fin k → m, ∀ g : Fin k → n, f.Injective → g.Injective → (A.submatrix f g).det = 0 ∨ From 06553ac2fd0df12b00670ebf31859d88c0bb9a52 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Thu, 21 Nov 2024 16:01:19 +0000 Subject: [PATCH 184/829] chore: lighten tactics (#19318) Change some tactic uses to more lightweight ones (but which still automate just as much): `nlinarith` to `positivity`/`linarith`/`gcongr`, `linarith` to `norm_num`. --- Mathlib/Algebra/Order/Floor.lean | 2 +- Mathlib/Analysis/Convex/Slope.lean | 4 ++-- Mathlib/Analysis/PSeries.lean | 3 +-- Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean | 2 +- Mathlib/NumberTheory/FermatPsp.lean | 2 +- Mathlib/RingTheory/Ideal/Operations.lean | 4 +--- Mathlib/Topology/MetricSpace/Kuratowski.lean | 10 ++++------ 7 files changed, 11 insertions(+), 16 deletions(-) diff --git a/Mathlib/Algebra/Order/Floor.lean b/Mathlib/Algebra/Order/Floor.lean index d3d6e7b16dc1d..0ec2dc284de96 100644 --- a/Mathlib/Algebra/Order/Floor.lean +++ b/Mathlib/Algebra/Order/Floor.lean @@ -1504,7 +1504,7 @@ theorem abs_sub_round_div_natCast_eq {m n : ℕ} : @[bound] theorem sub_half_lt_round (x : α) : x - 1 / 2 < round x := by - rw [round_eq x, show x - 1 / 2 = x + 1 / 2 - 1 by nlinarith] + rw [round_eq x, show x - 1 / 2 = x + 1 / 2 - 1 by linarith] exact Int.sub_one_lt_floor (x + 1 / 2) @[bound] diff --git a/Mathlib/Analysis/Convex/Slope.lean b/Mathlib/Analysis/Convex/Slope.lean index 1ecc6866a18aa..51a20a60734cd 100644 --- a/Mathlib/Analysis/Convex/Slope.lean +++ b/Mathlib/Analysis/Convex/Slope.lean @@ -40,7 +40,7 @@ theorem ConvexOn.slope_mono_adjacent (hf : ConvexOn 𝕜 s f) {x y z : 𝕜} (hx field_simp [a, b, mul_comm (z - x) _] at key ⊢ rw [div_le_div_iff_of_pos_right] · linarith - · nlinarith + · positivity /-- If `f : 𝕜 → 𝕜` is concave, then for any three points `x < y < z` the slope of the secant line of `f` on `[x, y]` is greater than the slope of the secant line of `f` on `[y, z]`. -/ @@ -73,7 +73,7 @@ theorem StrictConvexOn.slope_strict_mono_adjacent (hf : StrictConvexOn 𝕜 s f) field_simp [mul_comm (z - x) _] at key ⊢ rw [div_lt_div_iff_of_pos_right] · linarith - · nlinarith + · positivity /-- If `f : 𝕜 → 𝕜` is strictly concave, then for any three points `x < y < z` the slope of the secant line of `f` on `[x, y]` is strictly greater than the slope of the secant line of `f` on diff --git a/Mathlib/Analysis/PSeries.lean b/Mathlib/Analysis/PSeries.lean index 93a65386cf5e2..38fe34d857d76 100644 --- a/Mathlib/Analysis/PSeries.lean +++ b/Mathlib/Analysis/PSeries.lean @@ -389,8 +389,7 @@ theorem sum_Ioc_inv_sq_le_sub {k n : ℕ} (hk : k ≠ 0) (h : k ≤ n) : field_simp rw [div_le_div_iff₀ _ A] · linarith - · -- Porting note: was `nlinarith` - positivity + · positivity theorem sum_Ioo_inv_sq_le (k n : ℕ) : (∑ i ∈ Ioo k n, (i ^ 2 : α)⁻¹) ≤ 2 / (k + 1) := calc diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean index d05eed644c71c..71ae3fbd4d054 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean @@ -589,7 +589,7 @@ open Real Filter theorem tendsto_one_plus_div_rpow_exp (t : ℝ) : Tendsto (fun x : ℝ => (1 + t / x) ^ x) atTop (𝓝 (exp t)) := by apply ((Real.continuous_exp.tendsto _).comp (tendsto_mul_log_one_plus_div_atTop t)).congr' _ - have h₁ : (1 : ℝ) / 2 < 1 := by linarith + have h₁ : (1 : ℝ) / 2 < 1 := by norm_num have h₂ : Tendsto (fun x : ℝ => 1 + t / x) atTop (𝓝 1) := by simpa using (tendsto_inv_atTop_zero.const_mul t).const_add 1 refine (eventually_ge_of_tendsto_gt h₁ h₂).mono fun x hx => ?_ diff --git a/Mathlib/NumberTheory/FermatPsp.lean b/Mathlib/NumberTheory/FermatPsp.lean index b5c6b98a175d6..31cdde409beb6 100644 --- a/Mathlib/NumberTheory/FermatPsp.lean +++ b/Mathlib/NumberTheory/FermatPsp.lean @@ -304,7 +304,7 @@ private theorem psp_from_prime_gt_p {b : ℕ} (b_ge_two : 2 ≤ b) {p : ℕ} (p_ suffices h : p * b ^ 2 < (b ^ 2) ^ (p - 1) * b ^ 2 by apply gt_of_ge_of_gt · exact tsub_le_tsub_left (one_le_of_lt p_gt_two) ((b ^ 2) ^ (p - 1) * b ^ 2) - · have : p ≤ p * b ^ 2 := Nat.le_mul_of_pos_right _ (show 0 < b ^ 2 by nlinarith) + · have : p ≤ p * b ^ 2 := Nat.le_mul_of_pos_right _ (show 0 < b ^ 2 by positivity) exact tsub_lt_tsub_right_of_le this h suffices h : p < (b ^ 2) ^ (p - 1) by have : 4 ≤ b ^ 2 := by nlinarith diff --git a/Mathlib/RingTheory/Ideal/Operations.lean b/Mathlib/RingTheory/Ideal/Operations.lean index e18052a266b00..13dfba254995d 100644 --- a/Mathlib/RingTheory/Ideal/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Operations.lean @@ -441,9 +441,7 @@ lemma sup_pow_add_le_pow_sup_pow {n m : ℕ} : (I ⊔ J) ^ (n + m) ≤ I ^ n ⊔ ((Ideal.pow_le_pow_right hn).trans le_sup_left))) · refine (Ideal.mul_le_right.trans (Ideal.mul_le_left.trans ((Ideal.pow_le_pow_right ?_).trans le_sup_right))) - simp only [Finset.mem_range, Nat.lt_succ] at hi - rw [Nat.le_sub_iff_add_le hi] - nlinarith + omega variable (I J K) diff --git a/Mathlib/Topology/MetricSpace/Kuratowski.lean b/Mathlib/Topology/MetricSpace/Kuratowski.lean index 7b7e6e47ae7f0..a0fa5566df6e3 100644 --- a/Mathlib/Topology/MetricSpace/Kuratowski.lean +++ b/Mathlib/Topology/MetricSpace/Kuratowski.lean @@ -70,12 +70,10 @@ theorem embeddingOfSubset_isometry (H : DenseRange x) : Isometry (embeddingOfSub rw [C] gcongr _ ≤ 2 * (e / 2) + dist (embeddingOfSubset x b) (embeddingOfSubset x a) := by - have : |embeddingOfSubset x b n - embeddingOfSubset x a n| ≤ - dist (embeddingOfSubset x b) (embeddingOfSubset x a) := by - simp only [dist_eq_norm] - exact lp.norm_apply_le_norm ENNReal.top_ne_zero - (embeddingOfSubset x b - embeddingOfSubset x a) n - nlinarith + gcongr + simp only [dist_eq_norm] + exact lp.norm_apply_le_norm ENNReal.top_ne_zero + (embeddingOfSubset x b - embeddingOfSubset x a) n _ = dist (embeddingOfSubset x b) (embeddingOfSubset x a) + e := by ring simpa [dist_comm] using this From 2d53f5f766c9ed4438ad8327e84af7d714069f32 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Thu, 21 Nov 2024 17:01:14 +0000 Subject: [PATCH 185/829] =?UTF-8?q?feat(Data/Nat/Factorization/PrimePow):?= =?UTF-8?q?=20add=20equivalence=20Primes=20=C3=97=20=E2=84=95=20=E2=89=83?= =?UTF-8?q?=20PrimePowers=20(#19335)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This is part of the series of PRs leading to Dirichlet's Theorem on primes in AP. See [here](https://leanprover.zulipchat.com/#narrow/channel/144837-PR-reviews/topic/Prerequisites.20for.20PNT.20and.20Dirichlet's.20Thm/near/483303184) on Zulip. --- Mathlib/Data/Nat/Factorization/PrimePow.lean | 38 ++++++++++++++++++++ 1 file changed, 38 insertions(+) diff --git a/Mathlib/Data/Nat/Factorization/PrimePow.lean b/Mathlib/Data/Nat/Factorization/PrimePow.lean index 3d33d3f1dbd08..be263eea0756e 100644 --- a/Mathlib/Data/Nat/Factorization/PrimePow.lean +++ b/Mathlib/Data/Nat/Factorization/PrimePow.lean @@ -145,3 +145,41 @@ theorem Nat.mul_divisors_filter_prime_pow {a b : ℕ} (hab : a.Coprime b) : simp only [ha, hb, Finset.mem_union, Finset.mem_filter, Nat.mul_eq_zero, and_true, Ne, and_congr_left_iff, not_false_iff, Nat.mem_divisors, or_self_iff] apply hab.isPrimePow_dvd_mul + +lemma IsPrimePow.factorization_minFac_ne_zero {n : ℕ} (hn : IsPrimePow n) : + n.factorization n.minFac ≠ 0 := by + refine mt (Nat.factorization_eq_zero_iff _ _).mp ?_ + push_neg + exact ⟨n.minFac_prime hn.ne_one, n.minFac_dvd, hn.ne_zero⟩ + +/-- The canonical equivalence between pairs `(p, k)` with `p` a prime and `k : ℕ` +and the set of prime powers given by `(p, k) ↦ p^(k+1)`. -/ +def Nat.Primes.prodNatEquiv : Nat.Primes × ℕ ≃ {n : ℕ // IsPrimePow n} where + toFun pk := + ⟨pk.1 ^ (pk.2 + 1), ⟨pk.1, pk.2 + 1, prime_iff.mp pk.1.prop, pk.2.add_one_pos, rfl⟩⟩ + invFun n := + (⟨n.val.minFac, minFac_prime n.prop.ne_one⟩, n.val.factorization n.val.minFac - 1) + left_inv := fun (p, k) ↦ by + simp only [p.prop.pow_minFac k.add_one_ne_zero, Subtype.coe_eta, factorization_pow, p.prop, + Prime.factorization, Finsupp.smul_single, smul_eq_mul, mul_one, Finsupp.single_add, + Finsupp.coe_add, Pi.add_apply, Finsupp.single_eq_same, add_tsub_cancel_right] + right_inv n := by + ext1 + dsimp only + rw [sub_one_add_one n.prop.factorization_minFac_ne_zero, n.prop.minFac_pow_factorization_eq] + +@[simp] +lemma Nat.Primes.prodNatEquiv_apply (p : Nat.Primes) (k : ℕ) : + prodNatEquiv (p, k) = ⟨p ^ (k + 1), p, k + 1, prime_iff.mp p.prop, k.add_one_pos, rfl⟩ := by + rfl + +@[simp] +lemma Nat.Primes.coe_prodNatEquiv_apply (p : Nat.Primes) (k : ℕ) : + (prodNatEquiv (p, k) : ℕ) = p ^ (k + 1) := + rfl + +@[simp] +lemma Nat.Primes.prodNatEquiv_symm_apply {n : ℕ} (hn : IsPrimePow n) : + prodNatEquiv.symm ⟨n, hn⟩ = + (⟨n.minFac, minFac_prime hn.ne_one⟩, n.factorization n.minFac - 1) := + rfl From 439565f1bd4b2273a6cdcb4f4d31004ccfcb43f1 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 21 Nov 2024 21:37:19 +0000 Subject: [PATCH 186/829] refactor: use SignType to define IsTotallyUnimodular (#19345) This is a little easier than working with a three-element disjunction. [Zulip thread](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/Matrix.2EfromRows.20TU.20.28help.20wanted.29/near/483771954) --- .../Matrix/Determinant/TotallyUnimodular.lean | 26 +++++++++---------- 1 file changed, 12 insertions(+), 14 deletions(-) diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean index 3c81f42c1c1b9..cd7b6cae09bfd 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean @@ -5,6 +5,7 @@ Authors: Martin Dvorak, Vladimir Kolmogorov, Ivan Sergeev -/ import Mathlib.LinearAlgebra.Matrix.Determinant.Basic import Mathlib.Data.Matrix.ColumnRowPartitioned +import Mathlib.Data.Sign /-! # Totally unimodular matrices @@ -29,24 +30,20 @@ namespace Matrix variable {m n R : Type*} [CommRing R] /-- `A.IsTotallyUnimodular` means that every square submatrix of `A` (not necessarily contiguous) -has determinant `0` or `1` or `-1`. -/ +has determinant `0` or `1` or `-1`; that is, the determinant is in the range of `SignType.cast`. -/ def IsTotallyUnimodular (A : Matrix m n R) : Prop := ∀ k : ℕ, ∀ f : Fin k → m, ∀ g : Fin k → n, f.Injective → g.Injective → - (A.submatrix f g).det = 0 ∨ - (A.submatrix f g).det = 1 ∨ - (A.submatrix f g).det = -1 + (A.submatrix f g).det ∈ Set.range SignType.cast lemma isTotallyUnimodular_iff (A : Matrix m n R) : A.IsTotallyUnimodular ↔ ∀ k : ℕ, ∀ f : Fin k → m, ∀ g : Fin k → n, - (A.submatrix f g).det = 0 ∨ - (A.submatrix f g).det = 1 ∨ - (A.submatrix f g).det = -1 := by + (A.submatrix f g).det ∈ Set.range SignType.cast := by constructor <;> intro hA · intro k f g - if h : f.Injective ∧ g.Injective then - exact hA k f g h.1 h.2 - else - left + by_cases h : f.Injective ∧ g.Injective + · exact hA k f g h.1 h.2 + · refine ⟨0, ?_⟩ + rw [SignType.coe_zero, eq_comm] simp_rw [not_and_or, Function.not_injective_iff] at h obtain ⟨i, j, hfij, hij⟩ | ⟨i, j, hgij, hij⟩ := h · rw [← det_transpose, transpose_submatrix] @@ -59,10 +56,10 @@ lemma isTotallyUnimodular_iff (A : Matrix m n R) : A.IsTotallyUnimodular ↔ lemma IsTotallyUnimodular.apply {A : Matrix m n R} (hA : A.IsTotallyUnimodular) (i : m) (j : n) : - A i j = 0 ∨ A i j = 1 ∨ A i j = -1 := by + A i j ∈ Set.range SignType.cast := by let f : Fin 1 → m := (fun _ => i) let g : Fin 1 → n := (fun _ => j) - convert hA 1 f g (Function.injective_of_subsingleton f) (Function.injective_of_subsingleton g) <;> + convert hA 1 f g (Function.injective_of_subsingleton f) (Function.injective_of_subsingleton g) exact (det_fin_one (A.submatrix f g)).symm lemma IsTotallyUnimodular.submatrix {A : Matrix m n R} (hA : A.IsTotallyUnimodular) {k : ℕ} @@ -96,7 +93,8 @@ lemma fromRows_row0_isTotallyUnimodular_iff (A : Matrix m n R) {m' : Type*} : · exact hA k (Sum.inl ∘ f) g · if zerow : ∃ i, ∃ x', f i = Sum.inr x' then obtain ⟨i, _, _⟩ := zerow - left + use 0 + rw [eq_comm] apply det_eq_zero_of_row_eq_zero i intro simp_all From b51311315fdfbc2a8989b482a980f5277819425a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Martin=20Dvo=C5=99=C3=A1k?= Date: Thu, 21 Nov 2024 21:45:47 +0000 Subject: [PATCH 187/829] style(Data/Matrix/Block): define `Matrix.fromBlocks` more esthetically (#19324) --- Mathlib/Data/Matrix/Block.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Data/Matrix/Block.lean b/Mathlib/Data/Matrix/Block.lean index d572c642d2bca..cb4875d44825f 100644 --- a/Mathlib/Data/Matrix/Block.lean +++ b/Mathlib/Data/Matrix/Block.lean @@ -40,7 +40,7 @@ dimensions. -/ @[pp_nodot] def fromBlocks (A : Matrix n l α) (B : Matrix n m α) (C : Matrix o l α) (D : Matrix o m α) : Matrix (n ⊕ o) (l ⊕ m) α := - of <| Sum.elim (fun i => Sum.elim (A i) (B i)) fun i => Sum.elim (C i) (D i) + of <| Sum.elim (fun i => Sum.elim (A i) (B i)) (fun j => Sum.elim (C j) (D j)) @[simp] theorem fromBlocks_apply₁₁ (A : Matrix n l α) (B : Matrix n m α) (C : Matrix o l α) From ff67e337e215fb84ff403dd1ef4e99f51db6be8c Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 22 Nov 2024 00:49:48 +0000 Subject: [PATCH 188/829] chore(Topology/Order): move&generalize 4 lemmas (#19238) Generalize lemmas from `CompleteLinearOrder` to `ConditionallyCompleteLattice` or `ConditionallyCompleteLinearOrder`. --- .../MeasureTheory/Measure/Typeclasses.lean | 2 +- .../Topology/Algebra/Order/LiminfLimsup.lean | 37 ------------------- Mathlib/Topology/Order/OrderClosed.lean | 37 +++++++++++++++++++ 3 files changed, 38 insertions(+), 38 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/Typeclasses.lean b/Mathlib/MeasureTheory/Measure/Typeclasses.lean index 7ab48f4222914..c9a41869344f8 100644 --- a/Mathlib/MeasureTheory/Measure/Typeclasses.lean +++ b/Mathlib/MeasureTheory/Measure/Typeclasses.lean @@ -800,7 +800,7 @@ theorem countable_meas_pos_of_disjoint_of_meas_iUnion_ne_top₀ {ι : Type*} {_ simp only [fairmeas] rfl simpa only [fairmeas_eq, posmeas_def, ← preimage_iUnion, - iUnion_Ici_eq_Ioi_of_lt_of_tendsto (0 : ℝ≥0∞) (fun n => (as_mem n).1) as_lim] + iUnion_Ici_eq_Ioi_of_lt_of_tendsto (fun n => (as_mem n).1) as_lim] rw [countable_union] refine countable_iUnion fun n => Finite.countable ?_ exact finite_const_le_meas_of_disjoint_iUnion₀ μ (as_mem n).1 As_mble As_disj Union_As_finite diff --git a/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean b/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean index 44fe760b02479..bc4753090a8e3 100644 --- a/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean +++ b/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean @@ -411,43 +411,6 @@ theorem Monotone.map_liminf_of_continuousAt {f : R → S} (f_incr : Monotone f) end Monotone -section InfiAndSupr - -open Topology - -open Filter Set - -variable [CompleteLinearOrder R] [TopologicalSpace R] [OrderTopology R] - -theorem iInf_eq_of_forall_le_of_tendsto {x : R} {as : ι → R} (x_le : ∀ i, x ≤ as i) {F : Filter ι} - [Filter.NeBot F] (as_lim : Filter.Tendsto as F (𝓝 x)) : ⨅ i, as i = x := by - refine iInf_eq_of_forall_ge_of_forall_gt_exists_lt (fun i ↦ x_le i) ?_ - apply fun w x_lt_w ↦ ‹Filter.NeBot F›.nonempty_of_mem (eventually_lt_of_tendsto_lt x_lt_w as_lim) - -theorem iSup_eq_of_forall_le_of_tendsto {x : R} {as : ι → R} (le_x : ∀ i, as i ≤ x) {F : Filter ι} - [Filter.NeBot F] (as_lim : Filter.Tendsto as F (𝓝 x)) : ⨆ i, as i = x := - iInf_eq_of_forall_le_of_tendsto (R := Rᵒᵈ) le_x as_lim - -theorem iUnion_Ici_eq_Ioi_of_lt_of_tendsto (x : R) {as : ι → R} (x_lt : ∀ i, x < as i) - {F : Filter ι} [Filter.NeBot F] (as_lim : Filter.Tendsto as F (𝓝 x)) : - ⋃ i : ι, Ici (as i) = Ioi x := by - have obs : x ∉ range as := by - intro maybe_x_is - rcases mem_range.mp maybe_x_is with ⟨i, hi⟩ - simpa only [hi, lt_self_iff_false] using x_lt i - -- Porting note: `rw at *` was too destructive. Let's only rewrite `obs` and the goal. - have := iInf_eq_of_forall_le_of_tendsto (fun i ↦ (x_lt i).le) as_lim - rw [← this] at obs - rw [← this] - exact iUnion_Ici_eq_Ioi_iInf obs - -theorem iUnion_Iic_eq_Iio_of_lt_of_tendsto (x : R) {as : ι → R} (lt_x : ∀ i, as i < x) - {F : Filter ι} [Filter.NeBot F] (as_lim : Filter.Tendsto as F (𝓝 x)) : - ⋃ i : ι, Iic (as i) = Iio x := - iUnion_Ici_eq_Ioi_of_lt_of_tendsto (R := Rᵒᵈ) x lt_x as_lim - -end InfiAndSupr - section Indicator theorem limsup_eq_tendsto_sum_indicator_nat_atTop (s : ℕ → Set α) : diff --git a/Mathlib/Topology/Order/OrderClosed.lean b/Mathlib/Topology/Order/OrderClosed.lean index 07a4ecd8f0010..4efa005c67f3f 100644 --- a/Mathlib/Topology/Order/OrderClosed.lean +++ b/Mathlib/Topology/Order/OrderClosed.lean @@ -150,6 +150,10 @@ theorem disjoint_nhds_atBot_iff : Disjoint (𝓝 a) atBot ↔ ¬IsBot a := by refine disjoint_of_disjoint_of_mem disjoint_compl_left ?_ (Iic_mem_atBot b) exact isClosed_Iic.isOpen_compl.mem_nhds hb +theorem IsLUB.range_of_tendsto {F : Filter β} [F.NeBot] (hle : ∀ i, f i ≤ a) + (hlim : Tendsto f F (𝓝 a)) : IsLUB (range f) a := + ⟨forall_mem_range.mpr hle, fun _c hc ↦ le_of_tendsto' hlim fun i ↦ hc <| mem_range_self i⟩ + end Preorder section NoBotOrder @@ -170,6 +174,23 @@ theorem not_tendsto_atBot_of_tendsto_nhds (hf : Tendsto f l (𝓝 a)) : ¬Tendst end NoBotOrder +theorem iSup_eq_of_forall_le_of_tendsto {ι : Type*} {F : Filter ι} [Filter.NeBot F] + [ConditionallyCompleteLattice α] [TopologicalSpace α] [ClosedIicTopology α] + {a : α} {f : ι → α} (hle : ∀ i, f i ≤ a) (hlim : Filter.Tendsto f F (𝓝 a)) : + ⨆ i, f i = a := + have := F.nonempty_of_neBot + (IsLUB.range_of_tendsto hle hlim).ciSup_eq + +theorem iUnion_Iic_eq_Iio_of_lt_of_tendsto {ι : Type*} {F : Filter ι} [F.NeBot] + [ConditionallyCompleteLinearOrder α] [TopologicalSpace α] [ClosedIicTopology α] + {a : α} {f : ι → α} (hlt : ∀ i, f i < a) (hlim : Tendsto f F (𝓝 a)) : + ⋃ i : ι, Iic (f i) = Iio a := by + have obs : a ∉ range f := by + rw [mem_range] + rintro ⟨i, rfl⟩ + exact (hlt i).false + rw [← biUnion_range, (IsLUB.range_of_tendsto (le_of_lt <| hlt ·) hlim).biUnion_Iic_eq_Iio obs] + section LinearOrder variable [TopologicalSpace α] [LinearOrder α] [ClosedIicTopology α] [TopologicalSpace β] @@ -376,6 +397,10 @@ protected alias ⟨_, BddBelow.closure⟩ := bddBelow_closure theorem disjoint_nhds_atTop_iff : Disjoint (𝓝 a) atTop ↔ ¬IsTop a := disjoint_nhds_atBot_iff (α := αᵒᵈ) +theorem IsGLB.range_of_tendsto {F : Filter β} [F.NeBot] (hle : ∀ i, a ≤ f i) + (hlim : Tendsto f F (𝓝 a)) : IsGLB (range f) a := + IsLUB.range_of_tendsto (α := αᵒᵈ) hle hlim + end Preorder section NoTopOrder @@ -396,6 +421,18 @@ theorem not_tendsto_atTop_of_tendsto_nhds (hf : Tendsto f l (𝓝 a)) : ¬Tendst end NoTopOrder +theorem iInf_eq_of_forall_le_of_tendsto {ι : Type*} {F : Filter ι} [F.NeBot] + [ConditionallyCompleteLattice α] [TopologicalSpace α] [ClosedIciTopology α] + {a : α} {f : ι → α} (hle : ∀ i, a ≤ f i) (hlim : Tendsto f F (𝓝 a)) : + ⨅ i, f i = a := + iSup_eq_of_forall_le_of_tendsto (α := αᵒᵈ) hle hlim + +theorem iUnion_Ici_eq_Ioi_of_lt_of_tendsto {ι : Type*} {F : Filter ι} [F.NeBot] + [ConditionallyCompleteLinearOrder α] [TopologicalSpace α] [ClosedIciTopology α] + {a : α} {f : ι → α} (hlt : ∀ i, a < f i) (hlim : Tendsto f F (𝓝 a)) : + ⋃ i : ι, Ici (f i) = Ioi a := + iUnion_Iic_eq_Iio_of_lt_of_tendsto (α := αᵒᵈ) hlt hlim + section LinearOrder variable [TopologicalSpace α] [LinearOrder α] [ClosedIciTopology α] [TopologicalSpace β] From 1f86e43e4838ee944676f80cab98f4c66c3ed0a9 Mon Sep 17 00:00:00 2001 From: grunweg Date: Fri, 22 Nov 2024 01:48:35 +0000 Subject: [PATCH 189/829] feat: rewrite the trailing whitespace linter in Lean (#16334) Co-authored-by: Michael Rothgang --- Mathlib/Tactic/Linter/TextBased.lean | 26 ++++++++++++++++++++------ scripts/lint-style.py | 15 +-------------- 2 files changed, 21 insertions(+), 20 deletions(-) diff --git a/Mathlib/Tactic/Linter/TextBased.lean b/Mathlib/Tactic/Linter/TextBased.lean index a1d3a9955b489..aa2c455269b43 100644 --- a/Mathlib/Tactic/Linter/TextBased.lean +++ b/Mathlib/Tactic/Linter/TextBased.lean @@ -56,6 +56,8 @@ inductive StyleError where | adaptationNote /-- A line ends with windows line endings (\r\n) instead of unix ones (\n). -/ | windowsLineEnding + /-- A line contains trailing whitespace. -/ + | trailingWhitespace deriving BEq /-- How to format style errors -/ @@ -76,6 +78,7 @@ def StyleError.errorMessage (err : StyleError) : String := match err with "Found the string \"Adaptation note:\", please use the #adaptation_note command instead" | windowsLineEnding => "This line ends with a windows line ending (\r\n): please use Unix line\ endings (\n) instead" + | trailingWhitespace => "This line ends with some whitespace: please remove this" /-- The error code for a given style error. Keep this in sync with `parse?_errorContext` below! -/ -- FUTURE: we're matching the old codes in `lint-style.py` for compatibility; @@ -83,6 +86,7 @@ def StyleError.errorMessage (err : StyleError) : String := match err with def StyleError.errorCode (err : StyleError) : String := match err with | StyleError.adaptationNote => "ERR_ADN" | StyleError.windowsLineEnding => "ERR_WIN" + | StyleError.trailingWhitespace => "ERR_TWS" /-- Context for a style error: the actual error, the line number in the file we're reading and the path to the file. -/ @@ -160,6 +164,7 @@ def parse?_errorContext (line : String) : Option ErrorContext := Id.run do -- Use default values for parameters which are ignored for comparing style exceptions. -- NB: keep this in sync with `compare` above! | "ERR_ADN" => some (StyleError.adaptationNote) + | "ERR_TWS" => some (StyleError.trailingWhitespace) | "ERR_WIN" => some (StyleError.windowsLineEnding) | _ => none match String.toNat? lineNumber with @@ -196,14 +201,24 @@ section /-- Lint on any occurrences of the string "Adaptation note:" or variants thereof. -/ def adaptationNoteLinter : TextbasedLinter := fun lines ↦ Id.run do let mut errors := Array.mkEmpty 0 - let mut lineNumber := 1 - for line in lines do + for (line, idx) in lines.zipWithIndex do -- We make this shorter to catch "Adaptation note", "adaptation note" and a missing colon. if line.containsSubstr "daptation note" then - errors := errors.push (StyleError.adaptationNote, lineNumber) - lineNumber := lineNumber + 1 + errors := errors.push (StyleError.adaptationNote, idx + 1) return (errors, none) + +/-- Lint a collection of input strings if one of them contains trailing whitespace. -/ +def trailingWhitespaceLinter : TextbasedLinter := fun lines ↦ Id.run do + let mut errors := Array.mkEmpty 0 + let mut fixedLines := lines + for (line, idx) in lines.zipWithIndex do + if line.back == ' ' then + errors := errors.push (StyleError.trailingWhitespace, idx + 1) + fixedLines := fixedLines.set! idx line.trimRight + return (errors, if errors.size > 0 then some fixedLines else none) + + /-- Whether a collection of lines consists *only* of imports, blank lines and single-line comments. In practice, this means it's an imports-only file and exempt from almost all linting. -/ def isImportsOnlyFile (lines : Array String) : Bool := @@ -215,7 +230,7 @@ end /-- All text-based linters registered in this file. -/ def allLinters : Array TextbasedLinter := #[ - adaptationNoteLinter + adaptationNoteLinter, trailingWhitespaceLinter ] @@ -257,7 +272,6 @@ def lintFile (path : FilePath) (exceptions : Array ErrorContext) : (allOutput.flatten.filter (fun e ↦ (e.find?_comparable exceptions).isNone)) return (errors, if changes_made then some changed else none) - /-- Lint a collection of modules for style violations. Print formatted errors for all unexpected style violations to standard output; correct automatically fixable style errors if configured so. diff --git a/scripts/lint-style.py b/scripts/lint-style.py index f71277d979f42..9e6353bfc09ab 100755 --- a/scripts/lint-style.py +++ b/scripts/lint-style.py @@ -39,7 +39,6 @@ ERR_IBY = 11 # isolated by ERR_IWH = 22 # isolated where ERR_SEM = 13 # the substring " ;" -ERR_TWS = 15 # trailing whitespace ERR_CLN = 16 # line starts with a colon ERR_IND = 17 # second line not correctly indented ERR_ARR = 18 # space after "←" @@ -107,15 +106,6 @@ def annotate_strings(enumerate_lines): continue yield line_nr, line, *rem, False -def line_endings_check(lines, path): - errors = [] - newlines = [] - for line_nr, line in lines: - if line.endswith(" \n"): - errors += [(ERR_TWS, line_nr, path)] - line = line.rstrip() + "\n" - newlines.append((line_nr, line)) - return errors, newlines def four_spaces_in_second_line(lines, path): # TODO: also fix the space for all lines before ":=", right now we only fix the line after @@ -248,8 +238,6 @@ def format_errors(errors): output_message(path, line_nr, "ERR_IWH", "Line is an isolated where") if errno == ERR_SEM: output_message(path, line_nr, "ERR_SEM", "Line contains a space before a semicolon") - if errno == ERR_TWS: - output_message(path, line_nr, "ERR_TWS", "Trailing whitespace detected on line") if errno == ERR_CLN: output_message(path, line_nr, "ERR_CLN", "Put : and := before line breaks, not after") if errno == ERR_IND: @@ -267,8 +255,7 @@ def lint(path, fix=False): lines = f.readlines() enum_lines = enumerate(lines, 1) newlines = enum_lines - for error_check in [line_endings_check, - four_spaces_in_second_line, + for error_check in [four_spaces_in_second_line, isolated_by_dot_semicolon_check, left_arrow_check, nonterminal_simp_check]: From 43c93172bdc359f844e718f6c24c5e958496c4ea Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 22 Nov 2024 02:00:10 +0000 Subject: [PATCH 190/829] chore(Topology/Order): rename 2 lemmas (#19261) ... to allow dot notation. --- Mathlib/Algebra/Order/Floor/Prime.lean | 2 +- Mathlib/Analysis/Asymptotics/TVS.lean | 2 +- Mathlib/Analysis/BoundedVariation.lean | 2 +- Mathlib/Analysis/Normed/Lp/lpSpace.lean | 2 +- .../Analysis/SpecialFunctions/Pow/Deriv.lean | 2 +- Mathlib/Analysis/SpecificLimits/Normed.lean | 6 +++--- .../Computability/AkraBazzi/AkraBazzi.lean | 4 ++-- Mathlib/Data/Real/Pi/Irrational.lean | 4 ++-- .../Function/AEEqOfIntegral.lean | 6 +++--- Mathlib/MeasureTheory/Function/LpSpace.lean | 3 +-- .../Integral/IntegralEqImproper.lean | 4 +--- .../MeasureTheory/Measure/Portmanteau.lean | 6 +++--- .../MeasureTheory/Measure/WithDensity.lean | 2 +- Mathlib/Topology/Order/IntermediateValue.lean | 17 ++++++++-------- Mathlib/Topology/Order/OrderClosed.lean | 20 +++++++++++++++---- 15 files changed, 45 insertions(+), 37 deletions(-) diff --git a/Mathlib/Algebra/Order/Floor/Prime.lean b/Mathlib/Algebra/Order/Floor/Prime.lean index 506029de21811..8158125cc8e49 100644 --- a/Mathlib/Algebra/Order/Floor/Prime.lean +++ b/Mathlib/Algebra/Order/Floor/Prime.lean @@ -39,7 +39,7 @@ theorem exists_prime_mul_pow_div_factorial_lt_one [LinearOrderedField K] [FloorR letI := Preorder.topology K haveI : OrderTopology K := ⟨rfl⟩ ((Filter.frequently_atTop.mpr Nat.exists_infinite_primes).and_eventually - (eventually_lt_of_tendsto_lt zero_lt_one + (Filter.Tendsto.eventually_lt_const zero_lt_one (FloorSemiring.tendsto_mul_pow_div_factorial_sub_atTop a c 1))).forall_exists_of_atTop (n + 1) diff --git a/Mathlib/Analysis/Asymptotics/TVS.lean b/Mathlib/Analysis/Asymptotics/TVS.lean index f116d2b07113c..e4d49370e6a0a 100644 --- a/Mathlib/Analysis/Asymptotics/TVS.lean +++ b/Mathlib/Analysis/Asymptotics/TVS.lean @@ -125,7 +125,7 @@ lemma isLittleOTVS_one [ContinuousSMul 𝕜 E] {f : α → E} {l : Filter α} : rcases NormedField.exists_one_lt_norm 𝕜 with ⟨c, hc⟩ obtain ⟨ε, hε₀, hε⟩ : ∃ ε : ℝ≥0, 0 < ε ∧ (ε * ‖c‖₊ / r : ℝ≥0∞) < 1 := by apply Eventually.exists_gt - refine eventually_lt_of_tendsto_lt zero_lt_one <| Continuous.tendsto' ?_ _ _ (by simp) + refine Continuous.tendsto' ?_ _ _ (by simp) |>.eventually_lt_const zero_lt_one fun_prop (disch := intros; first | apply ENNReal.coe_ne_top | positivity) filter_upwards [hr ε hε₀.ne'] with x hx refine mem_of_egauge_lt_one hUb (hx.trans_lt ?_) diff --git a/Mathlib/Analysis/BoundedVariation.lean b/Mathlib/Analysis/BoundedVariation.lean index 4f281d9bfca6d..4a2fafb5a8000 100644 --- a/Mathlib/Analysis/BoundedVariation.lean +++ b/Mathlib/Analysis/BoundedVariation.lean @@ -180,7 +180,7 @@ theorem lowerSemicontinuous_aux {ι : Type*} {F : ι → α → E} {p : Filter (𝓝 (∑ i ∈ Finset.range n, edist (f (u (i + 1))) (f (u i)))) := by apply tendsto_finset_sum exact fun i _ => Tendsto.edist (Ffs (u i.succ) (us i.succ)) (Ffs (u i) (us i)) - exact (eventually_gt_of_tendsto_gt hlt this).mono fun i h => h.trans_le (sum_le (F i) n um us) + exact (this.eventually_const_lt hlt).mono fun i h => h.trans_le (sum_le (F i) n um us) /-- The map `(eVariationOn · s)` is lower semicontinuous for pointwise convergence *on `s`*. Pointwise convergence on `s` is encoded here as uniform convergence on the family consisting of the diff --git a/Mathlib/Analysis/Normed/Lp/lpSpace.lean b/Mathlib/Analysis/Normed/Lp/lpSpace.lean index 4797943bc7ec5..b2a36ecac89b5 100644 --- a/Mathlib/Analysis/Normed/Lp/lpSpace.lean +++ b/Mathlib/Analysis/Normed/Lp/lpSpace.lean @@ -183,7 +183,7 @@ theorem of_exponent_ge {p q : ℝ≥0∞} {f : ∀ i, E i} (hfq : Memℓp f q) ( have hf' := hfq.summable hq refine .of_norm_bounded_eventually _ hf' (@Set.Finite.subset _ { i | 1 ≤ ‖f i‖ } ?_ _ ?_) · have H : { x : α | 1 ≤ ‖f x‖ ^ q.toReal }.Finite := by - simpa using eventually_lt_of_tendsto_lt (by norm_num) hf'.tendsto_cofinite_zero + simpa using hf'.tendsto_cofinite_zero.eventually_lt_const (by norm_num) exact H.subset fun i hi => Real.one_le_rpow hi hq.le · show ∀ i, ¬|‖f i‖ ^ p.toReal| ≤ ‖f i‖ ^ q.toReal → 1 ≤ ‖f i‖ intro i hi diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean index 71ae3fbd4d054..49c7cb8a86df9 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean @@ -592,7 +592,7 @@ theorem tendsto_one_plus_div_rpow_exp (t : ℝ) : have h₁ : (1 : ℝ) / 2 < 1 := by norm_num have h₂ : Tendsto (fun x : ℝ => 1 + t / x) atTop (𝓝 1) := by simpa using (tendsto_inv_atTop_zero.const_mul t).const_add 1 - refine (eventually_ge_of_tendsto_gt h₁ h₂).mono fun x hx => ?_ + refine (h₂.eventually_const_le h₁).mono fun x hx => ?_ have hx' : 0 < 1 + t / x := by linarith simp [mul_comm x, exp_mul, exp_log hx'] diff --git a/Mathlib/Analysis/SpecificLimits/Normed.lean b/Mathlib/Analysis/SpecificLimits/Normed.lean index 69992fdb4a113..cf5f9286f5bd2 100644 --- a/Mathlib/Analysis/SpecificLimits/Normed.lean +++ b/Mathlib/Analysis/SpecificLimits/Normed.lean @@ -601,7 +601,7 @@ theorem summable_of_ratio_test_tendsto_lt_one {α : Type*} [NormedAddCommGroup (h : Tendsto (fun n ↦ ‖f (n + 1)‖ / ‖f n‖) atTop (𝓝 l)) : Summable f := by rcases exists_between hl₁ with ⟨r, hr₀, hr₁⟩ refine summable_of_ratio_norm_eventually_le hr₁ ?_ - filter_upwards [eventually_le_of_tendsto_lt hr₀ h, hf] with _ _ h₁ + filter_upwards [h.eventually_le_const hr₀, hf] with _ _ h₁ rwa [← div_le_iff₀ (norm_pos_iff.mpr h₁)] theorem not_summable_of_ratio_norm_eventually_ge {α : Type*} [SeminormedAddCommGroup α] {f : ℕ → α} @@ -629,12 +629,12 @@ theorem not_summable_of_ratio_test_tendsto_gt_one {α : Type*} [SeminormedAddCom {f : ℕ → α} {l : ℝ} (hl : 1 < l) (h : Tendsto (fun n ↦ ‖f (n + 1)‖ / ‖f n‖) atTop (𝓝 l)) : ¬Summable f := by have key : ∀ᶠ n in atTop, ‖f n‖ ≠ 0 := by - filter_upwards [eventually_ge_of_tendsto_gt hl h] with _ hn hc + filter_upwards [h.eventually_const_le hl] with _ hn hc rw [hc, _root_.div_zero] at hn linarith rcases exists_between hl with ⟨r, hr₀, hr₁⟩ refine not_summable_of_ratio_norm_eventually_ge hr₀ key.frequently ?_ - filter_upwards [eventually_ge_of_tendsto_gt hr₁ h, key] with _ _ h₁ + filter_upwards [h.eventually_const_le hr₁, key] with _ _ h₁ rwa [← le_div_iff₀ (lt_of_le_of_ne (norm_nonneg _) h₁.symm)] section NormedDivisionRing diff --git a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean index 9cfbcd713b62a..7cd217fadb8cf 100644 --- a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean +++ b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean @@ -544,9 +544,9 @@ lemma tendsto_atTop_sumCoeffsExp : Tendsto (fun (p : ℝ) => ∑ i, a i * (b i) lemma one_mem_range_sumCoeffsExp : 1 ∈ Set.range (fun (p : ℝ) => ∑ i, a i * (b i) ^ p) := by refine mem_range_of_exists_le_of_exists_ge R.continuous_sumCoeffsExp ?le_one ?ge_one case le_one => - exact Eventually.exists <| eventually_le_of_tendsto_lt zero_lt_one R.tendsto_zero_sumCoeffsExp + exact R.tendsto_zero_sumCoeffsExp.eventually_le_const zero_lt_one |>.exists case ge_one => - exact Eventually.exists <| R.tendsto_atTop_sumCoeffsExp.eventually_ge_atTop _ + exact R.tendsto_atTop_sumCoeffsExp.eventually_ge_atTop _ |>.exists /-- The function x ↦ ∑ a_i b_i^x is injective. This implies the uniqueness of `p`. -/ lemma injective_sumCoeffsExp : Function.Injective (fun (p : ℝ) => ∑ i, a i * (b i) ^ p) := diff --git a/Mathlib/Data/Real/Pi/Irrational.lean b/Mathlib/Data/Real/Pi/Irrational.lean index e3884b774698d..afebb88a105b2 100644 --- a/Mathlib/Data/Real/Pi/Irrational.lean +++ b/Mathlib/Data/Real/Pi/Irrational.lean @@ -283,8 +283,8 @@ private lemma not_irrational_exists_rep {x : ℝ} : 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) + have := (tendsto_pow_div_factorial_at_top_aux a).eventually_lt_const + (show (0 : ℝ) < 1 / 2 by norm_num) 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)) diff --git a/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean b/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean index ce393e78fea5c..10f195e4aacff 100644 --- a/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean +++ b/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean @@ -193,12 +193,12 @@ theorem ae_le_of_forall_setLIntegral_le_of_sigmaFinite₀ [SigmaFinite μ] have : Tendsto (fun n => g x + u n) atTop (𝓝 (g x + (0 : ℝ≥0))) := tendsto_const_nhds.add (ENNReal.tendsto_coe.2 u_lim) simp only [ENNReal.coe_zero, add_zero] at this - exact eventually_le_of_tendsto_lt hx this + exact this.eventually_le_const hx have L2 : ∀ᶠ n : ℕ in (atTop : Filter ℕ), g x ≤ (n : ℝ≥0) := - haveI : Tendsto (fun n : ℕ => ((n : ℝ≥0) : ℝ≥0∞)) atTop (𝓝 ∞) := by + have : Tendsto (fun n : ℕ => ((n : ℝ≥0) : ℝ≥0∞)) atTop (𝓝 ∞) := by simp only [ENNReal.coe_natCast] exact ENNReal.tendsto_nat_nhds_top - eventually_ge_of_tendsto_gt (hx.trans_le le_top) this + this.eventually_const_le (hx.trans_le le_top) apply Set.mem_iUnion.2 exact ((L1.and L2).and (eventually_mem_spanningSets μ x)).exists refine le_antisymm ?_ bot_le diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index 01ccd42815470..5b9c9290dac61 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -673,8 +673,7 @@ theorem exists_eLpNorm_indicator_le (hp : p ≠ ∞) (c : E) {ε : ℝ≥0∞} ( convert (NNReal.continuousAt_rpow_const (Or.inr hp₀')).tendsto.const_mul _ simp [hp₀''.ne'] have hε' : 0 < ε := hε.bot_lt - obtain ⟨δ, hδ, hδε'⟩ := - NNReal.nhds_zero_basis.eventually_iff.mp (eventually_le_of_tendsto_lt hε' this) + obtain ⟨δ, hδ, hδε'⟩ := NNReal.nhds_zero_basis.eventually_iff.mp (this.eventually_le_const hε') obtain ⟨η, hη, hηδ⟩ := exists_between hδ refine ⟨η, hη, ?_⟩ rw [← ENNReal.coe_rpow_of_nonneg _ hp₀', ← ENNReal.coe_mul] diff --git a/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean b/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean index e2ec1a0114d79..9678b171ff3fc 100644 --- a/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean +++ b/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean @@ -388,9 +388,7 @@ theorem AECover.iSup_lintegral_eq_of_countably_generated [Nonempty ι] [l.NeBot] have := hφ.lintegral_tendsto_of_countably_generated hfm refine ciSup_eq_of_forall_le_of_forall_lt_exists_gt (fun i => lintegral_mono' Measure.restrict_le_self le_rfl) fun w hw => ?_ - rcases exists_between hw with ⟨m, hm₁, hm₂⟩ - rcases (eventually_ge_of_tendsto_gt hm₂ this).exists with ⟨i, hi⟩ - exact ⟨i, lt_of_lt_of_le hm₁ hi⟩ + exact (this.eventually_const_lt hw).exists end Lintegral diff --git a/Mathlib/MeasureTheory/Measure/Portmanteau.lean b/Mathlib/MeasureTheory/Measure/Portmanteau.lean index 4d82e897d1771..72fe3305828be 100644 --- a/Mathlib/MeasureTheory/Measure/Portmanteau.lean +++ b/Mathlib/MeasureTheory/Measure/Portmanteau.lean @@ -278,14 +278,14 @@ theorem FiniteMeasure.limsup_measure_closed_le_of_tendsto {Ω ι : Type*} {L : F HasOuterApproxClosed.tendsto_lintegral_apprSeq F_closed (μ : Measure Ω) have room₁ : (μ : Measure Ω) F < (μ : Measure Ω) F + ε / 2 := ENNReal.lt_add_right (measure_lt_top (μ : Measure Ω) F).ne ε_pos' - obtain ⟨M, hM⟩ := eventually_atTop.mp <| eventually_lt_of_tendsto_lt room₁ key₁ + obtain ⟨M, hM⟩ := eventually_atTop.mp <| key₁.eventually_lt_const room₁ have key₂ := FiniteMeasure.tendsto_iff_forall_lintegral_tendsto.mp μs_lim (fs M) have room₂ : (lintegral (μ : Measure Ω) fun a ↦ fs M a) < (lintegral (μ : Measure Ω) fun a ↦ fs M a) + ε / 2 := ENNReal.lt_add_right (ne_of_lt ((fs M).lintegral_lt_top_of_nnreal _)) ε_pos' - have ev_near := Eventually.mono (eventually_lt_of_tendsto_lt room₂ key₂) fun n ↦ le_of_lt - have ev_near' := Eventually.mono ev_near + have ev_near := key₂.eventually_le_const room₂ + have ev_near' := ev_near.mono (fun n ↦ le_trans (HasOuterApproxClosed.measure_le_lintegral F_closed (μs n) M)) apply (Filter.limsup_le_limsup ev_near').trans rw [limsup_const] diff --git a/Mathlib/MeasureTheory/Measure/WithDensity.lean b/Mathlib/MeasureTheory/Measure/WithDensity.lean index d343307b3e3fc..f8c9285e8bbbf 100644 --- a/Mathlib/MeasureTheory/Measure/WithDensity.lean +++ b/Mathlib/MeasureTheory/Measure/WithDensity.lean @@ -669,7 +669,7 @@ lemma IsLocallyFiniteMeasure.withDensity_coe {f : α → ℝ≥0} (hf : Continuo IsLocallyFiniteMeasure (μ.withDensity fun x ↦ f x) := by refine ⟨fun x ↦ ?_⟩ rcases (μ.finiteAt_nhds x).exists_mem_basis ((nhds_basis_opens' x).restrict_subset - (eventually_le_of_tendsto_lt (lt_add_one _) (hf.tendsto x))) with ⟨U, ⟨⟨hUx, hUo⟩, hUf⟩, hμU⟩ + ((hf.tendsto x).eventually_le_const (lt_add_one _))) with ⟨U, ⟨⟨hUx, hUo⟩, hUf⟩, hμU⟩ refine ⟨U, hUx, ?_⟩ rw [withDensity_apply _ hUo.measurableSet] exact setLIntegral_lt_top_of_bddAbove hμU.ne ⟨f x + 1, forall_mem_image.2 hUf⟩ diff --git a/Mathlib/Topology/Order/IntermediateValue.lean b/Mathlib/Topology/Order/IntermediateValue.lean index 81446ea8fdf08..ef73d422348e1 100644 --- a/Mathlib/Topology/Order/IntermediateValue.lean +++ b/Mathlib/Topology/Order/IntermediateValue.lean @@ -122,21 +122,20 @@ theorem IsPreconnected.intermediate_value {s : Set X} (hs : IsPreconnected s) {a theorem IsPreconnected.intermediate_value_Ico {s : Set X} (hs : IsPreconnected s) {a : X} {l : Filter X} (ha : a ∈ s) [NeBot l] (hl : l ≤ 𝓟 s) {f : X → α} (hf : ContinuousOn f s) {v : α} (ht : Tendsto f l (𝓝 v)) : Ico (f a) v ⊆ f '' s := fun _ h => - hs.intermediate_value₂_eventually₁ ha hl hf continuousOn_const h.1 - (eventually_ge_of_tendsto_gt h.2 ht) + hs.intermediate_value₂_eventually₁ ha hl hf continuousOn_const h.1 (ht.eventually_const_le h.2) theorem IsPreconnected.intermediate_value_Ioc {s : Set X} (hs : IsPreconnected s) {a : X} {l : Filter X} (ha : a ∈ s) [NeBot l] (hl : l ≤ 𝓟 s) {f : X → α} (hf : ContinuousOn f s) {v : α} (ht : Tendsto f l (𝓝 v)) : Ioc v (f a) ⊆ f '' s := fun _ h => (hs.intermediate_value₂_eventually₁ ha hl continuousOn_const hf h.2 - (eventually_le_of_tendsto_lt h.1 ht)).imp fun _ h => h.imp_right Eq.symm + (ht.eventually_le_const h.1)).imp fun _ h => h.imp_right Eq.symm theorem IsPreconnected.intermediate_value_Ioo {s : Set X} (hs : IsPreconnected s) {l₁ l₂ : Filter X} [NeBot l₁] [NeBot l₂] (hl₁ : l₁ ≤ 𝓟 s) (hl₂ : l₂ ≤ 𝓟 s) {f : X → α} (hf : ContinuousOn f s) {v₁ v₂ : α} (ht₁ : Tendsto f l₁ (𝓝 v₁)) (ht₂ : Tendsto f l₂ (𝓝 v₂)) : Ioo v₁ v₂ ⊆ f '' s := fun _ h => hs.intermediate_value₂_eventually₂ hl₁ hl₂ hf continuousOn_const - (eventually_le_of_tendsto_lt h.1 ht₁) (eventually_ge_of_tendsto_gt h.2 ht₂) + (ht₁.eventually_le_const h.1) (ht₂.eventually_const_le h.2) theorem IsPreconnected.intermediate_value_Ici {s : Set X} (hs : IsPreconnected s) {a : X} {l : Filter X} (ha : a ∈ s) [NeBot l] (hl : l ≤ 𝓟 s) {f : X → α} (hf : ContinuousOn f s) @@ -153,19 +152,19 @@ theorem IsPreconnected.intermediate_value_Ioi {s : Set X} (hs : IsPreconnected s [NeBot l₁] [NeBot l₂] (hl₁ : l₁ ≤ 𝓟 s) (hl₂ : l₂ ≤ 𝓟 s) {f : X → α} (hf : ContinuousOn f s) {v : α} (ht₁ : Tendsto f l₁ (𝓝 v)) (ht₂ : Tendsto f l₂ atTop) : Ioi v ⊆ f '' s := fun y h => hs.intermediate_value₂_eventually₂ hl₁ hl₂ hf continuousOn_const - (eventually_le_of_tendsto_lt h ht₁) (tendsto_atTop.1 ht₂ y) + (ht₁.eventually_le_const h) (ht₂.eventually_ge_atTop y) theorem IsPreconnected.intermediate_value_Iio {s : Set X} (hs : IsPreconnected s) {l₁ l₂ : Filter X} [NeBot l₁] [NeBot l₂] (hl₁ : l₁ ≤ 𝓟 s) (hl₂ : l₂ ≤ 𝓟 s) {f : X → α} (hf : ContinuousOn f s) {v : α} (ht₁ : Tendsto f l₁ atBot) (ht₂ : Tendsto f l₂ (𝓝 v)) : Iio v ⊆ f '' s := fun y h => - hs.intermediate_value₂_eventually₂ hl₁ hl₂ hf continuousOn_const (tendsto_atBot.1 ht₁ y) - (eventually_ge_of_tendsto_gt h ht₂) + hs.intermediate_value₂_eventually₂ hl₁ hl₂ hf continuousOn_const (ht₁.eventually_le_atBot y) + (ht₂.eventually_const_le h) theorem IsPreconnected.intermediate_value_Iii {s : Set X} (hs : IsPreconnected s) {l₁ l₂ : Filter X} [NeBot l₁] [NeBot l₂] (hl₁ : l₁ ≤ 𝓟 s) (hl₂ : l₂ ≤ 𝓟 s) {f : X → α} (hf : ContinuousOn f s) (ht₁ : Tendsto f l₁ atBot) (ht₂ : Tendsto f l₂ atTop) : univ ⊆ f '' s := fun y _ => - hs.intermediate_value₂_eventually₂ hl₁ hl₂ hf continuousOn_const (tendsto_atBot.1 ht₁ y) - (tendsto_atTop.1 ht₂ y) + hs.intermediate_value₂_eventually₂ hl₁ hl₂ hf continuousOn_const (ht₁.eventually_le_atBot y) + (ht₂.eventually_ge_atTop y) /-- **Intermediate Value Theorem** for continuous functions on connected spaces. -/ theorem intermediate_value_univ [PreconnectedSpace X] (a b : X) {f : X → α} (hf : Continuous f) : diff --git a/Mathlib/Topology/Order/OrderClosed.lean b/Mathlib/Topology/Order/OrderClosed.lean index 4efa005c67f3f..7ded140f87a07 100644 --- a/Mathlib/Topology/Order/OrderClosed.lean +++ b/Mathlib/Topology/Order/OrderClosed.lean @@ -213,14 +213,20 @@ theorem Ici_mem_nhds (h : a < b) : Ici a ∈ 𝓝 b := theorem eventually_ge_nhds (hab : b < a) : ∀ᶠ x in 𝓝 a, b ≤ x := Ici_mem_nhds hab -theorem eventually_gt_of_tendsto_gt {l : Filter γ} {f : γ → α} {u v : α} (hv : u < v) +theorem Filter.Tendsto.eventually_const_lt {l : Filter γ} {f : γ → α} {u v : α} (hv : u < v) (h : Filter.Tendsto f l (𝓝 v)) : ∀ᶠ a in l, u < f a := h.eventually <| eventually_gt_nhds hv -theorem eventually_ge_of_tendsto_gt {l : Filter γ} {f : γ → α} {u v : α} (hv : u < v) +@[deprecated (since := "2024-11-17")] +alias eventually_gt_of_tendsto_gt := Filter.Tendsto.eventually_const_lt + +theorem Filter.Tendsto.eventually_const_le {l : Filter γ} {f : γ → α} {u v : α} (hv : u < v) (h : Tendsto f l (𝓝 v)) : ∀ᶠ a in l, u ≤ f a := h.eventually <| eventually_ge_nhds hv +@[deprecated (since := "2024-11-17")] +alias eventually_ge_of_tendsto_gt := Filter.Tendsto.eventually_const_le + protected theorem Dense.exists_gt [NoMaxOrder α] {s : Set α} (hs : Dense s) (x : α) : ∃ y ∈ s, x < y := hs.exists_mem_open isOpen_Ioi (exists_gt x) @@ -451,14 +457,20 @@ theorem Iic_mem_nhds (h : a < b) : Iic b ∈ 𝓝 a := theorem eventually_le_nhds (hab : a < b) : ∀ᶠ x in 𝓝 a, x ≤ b := Iic_mem_nhds hab -theorem eventually_lt_of_tendsto_lt {l : Filter γ} {f : γ → α} {u v : α} (hv : v < u) +theorem Filter.Tendsto.eventually_lt_const {l : Filter γ} {f : γ → α} {u v : α} (hv : v < u) (h : Filter.Tendsto f l (𝓝 v)) : ∀ᶠ a in l, f a < u := h.eventually <| eventually_lt_nhds hv -theorem eventually_le_of_tendsto_lt {l : Filter γ} {f : γ → α} {u v : α} (hv : v < u) +@[deprecated (since := "2024-11-17")] +alias eventually_lt_of_tendsto_lt := Filter.Tendsto.eventually_lt_const + +theorem Filter.Tendsto.eventually_le_const {l : Filter γ} {f : γ → α} {u v : α} (hv : v < u) (h : Tendsto f l (𝓝 v)) : ∀ᶠ a in l, f a ≤ u := h.eventually <| eventually_le_nhds hv +@[deprecated (since := "2024-11-17")] +alias eventually_le_of_tendsto_lt := Filter.Tendsto.eventually_le_const + protected theorem Dense.exists_lt [NoMinOrder α] {s : Set α} (hs : Dense s) (x : α) : ∃ y ∈ s, y < x := hs.orderDual.exists_gt x From fe0e8bcc2ddc54118bbd2abdc474dd1d2e3076f3 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Fri, 22 Nov 2024 04:40:29 +0000 Subject: [PATCH 191/829] feat(NumberTheory/LSeries/PrimesInAP): new file (#19344) This PR creates a new file `Mathlib.NumberTheory.LSeries.PrimesInAP` that will eventually contain a proof of **Dirichlet's Theorem on primes in arithmetic progression**. As a first step, we provide two lemmas on re-writing sums (or products) over a function supported in prime powers. (We also add an API lemma for `Multipliable`/`Summable` that seems to be missing.) See [here](https://leanprover.zulipchat.com/#narrow/channel/144837-PR-reviews/topic/Prerequisites.20for.20PNT.20and.20Dirichlet's.20Thm/near/483714489) on Zulip. --- Mathlib.lean | 1 + Mathlib/NumberTheory/LSeries/PrimesInAP.lean | 67 +++++++++++++++++++ .../Algebra/InfiniteSum/Constructions.lean | 5 ++ 3 files changed, 73 insertions(+) create mode 100644 Mathlib/NumberTheory/LSeries/PrimesInAP.lean diff --git a/Mathlib.lean b/Mathlib.lean index e6b62736993a2..5e0aee1763797 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3770,6 +3770,7 @@ import Mathlib.NumberTheory.LSeries.Linearity import Mathlib.NumberTheory.LSeries.MellinEqDirichlet import Mathlib.NumberTheory.LSeries.Nonvanishing import Mathlib.NumberTheory.LSeries.Positivity +import Mathlib.NumberTheory.LSeries.PrimesInAP import Mathlib.NumberTheory.LSeries.RiemannZeta import Mathlib.NumberTheory.LSeries.ZMod import Mathlib.NumberTheory.LegendreSymbol.AddCharacter diff --git a/Mathlib/NumberTheory/LSeries/PrimesInAP.lean b/Mathlib/NumberTheory/LSeries/PrimesInAP.lean new file mode 100644 index 0000000000000..0c1fc1a1ebf26 --- /dev/null +++ b/Mathlib/NumberTheory/LSeries/PrimesInAP.lean @@ -0,0 +1,67 @@ +/- +Copyright (c) 2024 Michael Stoll. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Michael Stoll +-/ +import Mathlib.Data.Nat.Factorization.PrimePow +import Mathlib.Topology.Algebra.InfiniteSum.Constructions +import Mathlib.Topology.Algebra.InfiniteSum.NatInt + +/-! +# Dirichlet's Theorem on primes in arithmetic progression + +The goal of this file is to prove **Dirichlet's Theorem**: If `q` is a positive natural number +and `a : ZMod q` is invertible, then there are infinitely many prime numbers `p` such that +`(p : ZMod q) = a`. + +This will be done in the following steps: + +1. Some auxiliary lemmas on infinite products and sums +2. (TODO) Results on the von Mangoldt function restricted to a residue class +3. (TODO) Results on its L-series +4. (TODO) Derivation of Dirichlet's Theorem +-/ + +/-! +### Auxiliary statements + +An infinite product or sum over a function supported in prime powers can be written +as an iterated product or sum over primes and natural numbers. +-/ + +section auxiliary + +variable {α β γ : Type*} [CommGroup α] [UniformSpace α] [UniformGroup α] [CompleteSpace α] + [T0Space α] + +open Nat.Primes in +@[to_additive tsum_eq_tsum_primes_of_support_subset_prime_powers] +lemma tprod_eq_tprod_primes_of_mulSupport_subset_prime_powers {f : ℕ → α} + (hfm : Multipliable f) (hf : Function.mulSupport f ⊆ {n | IsPrimePow n}) : + ∏' n : ℕ, f n = ∏' (p : Nat.Primes) (k : ℕ), f (p ^ (k + 1)) := by + have hfm' : Multipliable fun pk : Nat.Primes × ℕ ↦ f (pk.fst ^ (pk.snd + 1)) := + prodNatEquiv.symm.multipliable_iff.mp <| by + simpa only [← coe_prodNatEquiv_apply, Prod.eta, Function.comp_def, Equiv.apply_symm_apply] + using hfm.subtype _ + simp only [← tprod_subtype_eq_of_mulSupport_subset hf, Set.coe_setOf, ← prodNatEquiv.tprod_eq, + ← tprod_prod hfm'] + refine tprod_congr fun (p, k) ↦ congrArg f <| coe_prodNatEquiv_apply .. + +@[to_additive tsum_eq_tsum_primes_add_tsum_primes_of_support_subset_prime_powers] +lemma tprod_eq_tprod_primes_mul_tprod_primes_of_mulSupport_subset_prime_powers {f : ℕ → α} + (hfm : Multipliable f) (hf : Function.mulSupport f ⊆ {n | IsPrimePow n}) : + ∏' n : ℕ, f n = (∏' p : Nat.Primes, f p) * ∏' (p : Nat.Primes) (k : ℕ), f (p ^ (k + 2)) := by + rw [tprod_eq_tprod_primes_of_mulSupport_subset_prime_powers hfm hf] + have hfs' (p : Nat.Primes) : Multipliable fun k : ℕ ↦ f (p ^ (k + 1)) := + hfm.comp_injective <| (strictMono_nat_of_lt_succ + fun k ↦ pow_lt_pow_right₀ p.prop.one_lt <| lt_add_one (k + 1)).injective + conv_lhs => + enter [1, p]; rw [tprod_eq_zero_mul (hfs' p), zero_add, pow_one] + enter [2, 1, k]; rw [add_assoc, one_add_one_eq_two] + exact tprod_mul (Multipliable.subtype hfm _) <| + Multipliable.prod (f := fun (pk : Nat.Primes × ℕ) ↦ f (pk.1 ^ (pk.2 + 2))) <| + hfm.comp_injective <| Subtype.val_injective |>.comp + Nat.Primes.prodNatEquiv.injective |>.comp <| + Function.Injective.prodMap (fun ⦃_ _⦄ a ↦ a) <| add_left_injective 1 + +end auxiliary diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean b/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean index 47d09ed95c689..8b1f1a092440e 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean @@ -192,6 +192,11 @@ theorem Multipliable.prod_factor {f : β × γ → α} (h : Multipliable f) (b : Multipliable fun c ↦ f (b, c) := h.comp_injective fun _ _ h ↦ (Prod.ext_iff.1 h).2 +@[to_additive Summable.prod] +lemma Multipliable.prod {f : β × γ → α} (h : Multipliable f) : + Multipliable fun b ↦ ∏' c, f (b, c) := + ((Equiv.sigmaEquivProd β γ).multipliable_iff.mpr h).sigma + @[to_additive] lemma HasProd.tprod_fiberwise [T2Space α] {f : β → α} {a : α} (hf : HasProd f a) (g : β → γ) : HasProd (fun c : γ ↦ ∏' b : g ⁻¹' {c}, f b) a := From 981dca5c7e4eea030c1ea4e06efb682c3e784d9c Mon Sep 17 00:00:00 2001 From: sgouezel Date: Fri, 22 Nov 2024 05:01:05 +0000 Subject: [PATCH 192/829] chore(SchwartzSpace): clean up white space (#19336) Definitions in term mode using `mkCLM` are indented by 6 spaces, but formulating things with `refine` drops it to a more reasonable 2 spaces. Absolutely no math change, just unindenting. --- .../Analysis/Distribution/SchwartzSpace.lean | 234 +++++++++--------- 1 file changed, 114 insertions(+), 120 deletions(-) diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index 6debdf4d51000..0c1edc7031b70 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -764,53 +764,51 @@ variable [NormedAddCommGroup G] [NormedSpace ℝ G] /-- The map `f ↦ (x ↦ B (f x) (g x))` as a continuous `𝕜`-linear map on Schwartz space, where `B` is a continuous `𝕜`-linear map and `g` is a function of temperate growth. -/ def bilinLeftCLM (B : E →L[ℝ] F →L[ℝ] G) {g : D → F} (hg : g.HasTemperateGrowth) : - 𝓢(D, E) →L[ℝ] 𝓢(D, G) := + 𝓢(D, E) →L[ℝ] 𝓢(D, G) := by -- Todo (after port): generalize to `B : E →L[𝕜] F →L[𝕜] G` and `𝕜`-linear - mkCLM - (fun f x => B (f x) (g x)) + refine mkCLM (fun f x => B (f x) (g x)) (fun _ _ _ => by simp only [map_add, add_left_inj, Pi.add_apply, eq_self_iff_true, ContinuousLinearMap.add_apply]) (fun _ _ _ => by simp only [smul_apply, map_smul, ContinuousLinearMap.coe_smul', Pi.smul_apply, RingHom.id_apply]) - (fun f => (B.isBoundedBilinearMap.contDiff.restrict_scalars ℝ).comp (f.smooth'.prod hg.1)) - (by - rintro ⟨k, n⟩ - rcases hg.norm_iteratedFDeriv_le_uniform_aux n with ⟨l, C, hC, hgrowth⟩ - use - Finset.Iic (l + k, n), ‖B‖ * ((n : ℝ) + (1 : ℝ)) * n.choose (n / 2) * (C * 2 ^ (l + k)), - by positivity - intro f x - have hxk : 0 ≤ ‖x‖ ^ k := by positivity - have hnorm_mul := - ContinuousLinearMap.norm_iteratedFDeriv_le_of_bilinear B f.smooth' hg.1 x (n := n) le_top - refine le_trans (mul_le_mul_of_nonneg_left hnorm_mul hxk) ?_ - move_mul [← ‖B‖] - simp_rw [mul_assoc ‖B‖] - gcongr _ * ?_ - rw [Finset.mul_sum] - have : (∑ _x ∈ Finset.range (n + 1), (1 : ℝ)) = n + 1 := by simp - simp_rw [mul_assoc ((n : ℝ) + 1)] - rw [← this, Finset.sum_mul] - refine Finset.sum_le_sum fun i hi => ?_ - simp only [one_mul] - move_mul [(Nat.choose n i : ℝ), (Nat.choose n (n / 2) : ℝ)] - gcongr ?_ * ?_ - swap - · norm_cast - exact i.choose_le_middle n - specialize hgrowth (n - i) (by simp only [tsub_le_self]) x - refine le_trans (mul_le_mul_of_nonneg_left hgrowth (by positivity)) ?_ - move_mul [C] - gcongr ?_ * C - rw [Finset.mem_range_succ_iff] at hi - change i ≤ (l + k, n).snd at hi - refine le_trans ?_ (one_add_le_sup_seminorm_apply le_rfl hi f x) - rw [pow_add] - move_mul [(1 + ‖x‖) ^ l] - gcongr - simp) + (fun f => (B.isBoundedBilinearMap.contDiff.restrict_scalars ℝ).comp (f.smooth'.prod hg.1)) ?_ + rintro ⟨k, n⟩ + rcases hg.norm_iteratedFDeriv_le_uniform_aux n with ⟨l, C, hC, hgrowth⟩ + use + Finset.Iic (l + k, n), ‖B‖ * ((n : ℝ) + (1 : ℝ)) * n.choose (n / 2) * (C * 2 ^ (l + k)), + by positivity + intro f x + have hxk : 0 ≤ ‖x‖ ^ k := by positivity + have hnorm_mul := + ContinuousLinearMap.norm_iteratedFDeriv_le_of_bilinear B f.smooth' hg.1 x (n := n) le_top + refine le_trans (mul_le_mul_of_nonneg_left hnorm_mul hxk) ?_ + move_mul [← ‖B‖] + simp_rw [mul_assoc ‖B‖] + gcongr _ * ?_ + rw [Finset.mul_sum] + have : (∑ _x ∈ Finset.range (n + 1), (1 : ℝ)) = n + 1 := by simp + simp_rw [mul_assoc ((n : ℝ) + 1)] + rw [← this, Finset.sum_mul] + refine Finset.sum_le_sum fun i hi => ?_ + simp only [one_mul] + move_mul [(Nat.choose n i : ℝ), (Nat.choose n (n / 2) : ℝ)] + gcongr ?_ * ?_ + swap + · norm_cast + exact i.choose_le_middle n + specialize hgrowth (n - i) (by simp only [tsub_le_self]) x + refine le_trans (mul_le_mul_of_nonneg_left hgrowth (by positivity)) ?_ + move_mul [C] + gcongr ?_ * C + rw [Finset.mem_range_succ_iff] at hi + change i ≤ (l + k, n).snd at hi + refine le_trans ?_ (one_add_le_sup_seminorm_apply le_rfl hi f x) + rw [pow_add] + move_mul [(1 + ‖x‖) ^ l] + gcongr + simp end Multiplication @@ -824,68 +822,67 @@ variable [NormedSpace 𝕜 F] [SMulCommClass ℝ 𝕜 F] /-- Composition with a function on the right is a continuous linear map on Schwartz space provided that the function is temperate and growths polynomially near infinity. -/ def compCLM {g : D → E} (hg : g.HasTemperateGrowth) - (hg_upper : ∃ (k : ℕ) (C : ℝ), ∀ x, ‖x‖ ≤ C * (1 + ‖g x‖) ^ k) : 𝓢(E, F) →L[𝕜] 𝓢(D, F) := - mkCLM (fun f x => f (g x)) + (hg_upper : ∃ (k : ℕ) (C : ℝ), ∀ x, ‖x‖ ≤ C * (1 + ‖g x‖) ^ k) : 𝓢(E, F) →L[𝕜] 𝓢(D, F) := by + refine mkCLM (fun f x => f (g x)) (fun _ _ _ => by simp only [add_left_inj, Pi.add_apply, eq_self_iff_true]) (fun _ _ _ => rfl) - (fun f => f.smooth'.comp hg.1) - (by - rintro ⟨k, n⟩ - rcases hg.norm_iteratedFDeriv_le_uniform_aux n with ⟨l, C, hC, hgrowth⟩ - rcases hg_upper with ⟨kg, Cg, hg_upper'⟩ - have hCg : 1 ≤ 1 + Cg := by - refine le_add_of_nonneg_right ?_ - specialize hg_upper' 0 - rw [norm_zero] at hg_upper' - exact nonneg_of_mul_nonneg_left hg_upper' (by positivity) - let k' := kg * (k + l * n) - use Finset.Iic (k', n), (1 + Cg) ^ (k + l * n) * ((C + 1) ^ n * n ! * 2 ^ k'), by positivity - intro f x - let seminorm_f := ((Finset.Iic (k', n)).sup (schwartzSeminormFamily 𝕜 _ _)) f - have hg_upper'' : (1 + ‖x‖) ^ (k + l * n) ≤ (1 + Cg) ^ (k + l * n) * (1 + ‖g x‖) ^ k' := by - rw [pow_mul, ← mul_pow] - gcongr - rw [add_mul] - refine add_le_add ?_ (hg_upper' x) - nth_rw 1 [← one_mul (1 : ℝ)] - gcongr - 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 - intro i hi - have hpos : 0 < (1 + ‖g x‖) ^ k' := by positivity - rw [le_div_iff₀' hpos] - change i ≤ (k', n).snd at hi - exact one_add_le_sup_seminorm_apply le_rfl hi _ _ - have hgrowth' : ∀ N : ℕ, 1 ≤ N → N ≤ n → - ‖iteratedFDeriv ℝ N g x‖ ≤ ((C + 1) * (1 + ‖x‖) ^ l) ^ N := by - intro N hN₁ hN₂ - refine (hgrowth N hN₂ x).trans ?_ - rw [mul_pow] - have hN₁' := (lt_of_lt_of_le zero_lt_one hN₁).ne' - gcongr - · exact le_trans (by simp [hC]) (le_self_pow₀ (by simp [hC]) hN₁') - · refine le_self_pow₀ (one_le_pow₀ ?_) hN₁' - 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 := - pow_le_pow_left₀ (norm_nonneg _) (by simp only [zero_le_one, le_add_iff_nonneg_left]) _ - refine le_trans (mul_le_mul hxk this (by positivity) (by positivity)) ?_ - have rearrange : - (1 + ‖x‖) ^ k * - (n ! * (2 ^ k' * seminorm_f / (1 + ‖g x‖) ^ k') * ((C + 1) * (1 + ‖x‖) ^ l) ^ n) = - (1 + ‖x‖) ^ (k + l * n) / (1 + ‖g x‖) ^ k' * - ((C + 1) ^ n * n ! * 2 ^ k' * seminorm_f) := by - rw [mul_pow, pow_add, ← pow_mul] - ring - rw [rearrange] - have hgxk' : 0 < (1 + ‖g x‖) ^ k' := by positivity - rw [← div_le_iff₀ hgxk'] at hg_upper'' - have hpos : (0 : ℝ) ≤ (C + 1) ^ n * n ! * 2 ^ k' * seminorm_f := by - have : 0 ≤ seminorm_f := apply_nonneg _ _ - positivity - refine le_trans (mul_le_mul_of_nonneg_right hg_upper'' hpos) ?_ - rw [← mul_assoc]) + (fun f => f.smooth'.comp hg.1) ?_ + rintro ⟨k, n⟩ + rcases hg.norm_iteratedFDeriv_le_uniform_aux n with ⟨l, C, hC, hgrowth⟩ + rcases hg_upper with ⟨kg, Cg, hg_upper'⟩ + have hCg : 1 ≤ 1 + Cg := by + refine le_add_of_nonneg_right ?_ + specialize hg_upper' 0 + rw [norm_zero] at hg_upper' + exact nonneg_of_mul_nonneg_left hg_upper' (by positivity) + let k' := kg * (k + l * n) + use Finset.Iic (k', n), (1 + Cg) ^ (k + l * n) * ((C + 1) ^ n * n ! * 2 ^ k'), by positivity + intro f x + let seminorm_f := ((Finset.Iic (k', n)).sup (schwartzSeminormFamily 𝕜 _ _)) f + have hg_upper'' : (1 + ‖x‖) ^ (k + l * n) ≤ (1 + Cg) ^ (k + l * n) * (1 + ‖g x‖) ^ k' := by + rw [pow_mul, ← mul_pow] + gcongr + rw [add_mul] + refine add_le_add ?_ (hg_upper' x) + nth_rw 1 [← one_mul (1 : ℝ)] + gcongr + 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 + intro i hi + have hpos : 0 < (1 + ‖g x‖) ^ k' := by positivity + rw [le_div_iff₀' hpos] + change i ≤ (k', n).snd at hi + exact one_add_le_sup_seminorm_apply le_rfl hi _ _ + have hgrowth' : ∀ N : ℕ, 1 ≤ N → N ≤ n → + ‖iteratedFDeriv ℝ N g x‖ ≤ ((C + 1) * (1 + ‖x‖) ^ l) ^ N := by + intro N hN₁ hN₂ + refine (hgrowth N hN₂ x).trans ?_ + rw [mul_pow] + have hN₁' := (lt_of_lt_of_le zero_lt_one hN₁).ne' + gcongr + · exact le_trans (by simp [hC]) (le_self_pow₀ (by simp [hC]) hN₁') + · refine le_self_pow₀ (one_le_pow₀ ?_) hN₁' + 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 := + pow_le_pow_left₀ (norm_nonneg _) (by simp only [zero_le_one, le_add_iff_nonneg_left]) _ + refine le_trans (mul_le_mul hxk this (by positivity) (by positivity)) ?_ + have rearrange : + (1 + ‖x‖) ^ k * + (n ! * (2 ^ k' * seminorm_f / (1 + ‖g x‖) ^ k') * ((C + 1) * (1 + ‖x‖) ^ l) ^ n) = + (1 + ‖x‖) ^ (k + l * n) / (1 + ‖g x‖) ^ k' * + ((C + 1) ^ n * n ! * 2 ^ k' * seminorm_f) := by + rw [mul_pow, pow_add, ← pow_mul] + ring + rw [rearrange] + have hgxk' : 0 < (1 + ‖g x‖) ^ k' := by positivity + rw [← div_le_iff₀ hgxk'] at hg_upper'' + have hpos : (0 : ℝ) ≤ (C + 1) ^ n * n ! * 2 ^ k' * seminorm_f := by + have : 0 ≤ seminorm_f := apply_nonneg _ _ + positivity + refine le_trans (mul_le_mul_of_nonneg_right hg_upper'' hpos) ?_ + rw [← mul_assoc] @[simp] lemma compCLM_apply {g : D → E} (hg : g.HasTemperateGrowth) (hg_upper : ∃ (k : ℕ) (C : ℝ), ∀ x, ‖x‖ ≤ C * (1 + ‖g x‖) ^ k) (f : 𝓢(E, F)) : @@ -1070,25 +1067,22 @@ lemma integrable (f : 𝓢(D, V)) : Integrable f μ := variable (𝕜 μ) in /-- The integral as a continuous linear map from Schwartz space to the codomain. -/ -def integralCLM : 𝓢(D, V) →L[𝕜] V := - mkCLMtoNormedSpace (∫ x, · x ∂μ) - (fun f g ↦ by - exact integral_add f.integrable g.integrable) - (integral_smul · ·) - (by - rcases hμ.exists_integrable with ⟨n, h⟩ - let m := (n, 0) - use Finset.Iic m, 2 ^ n * ∫ x : D, (1 + ‖x‖) ^ (- (n : ℝ)) ∂μ - refine ⟨by positivity, fun f ↦ (norm_integral_le_integral_norm f).trans ?_⟩ - have h' : ∀ x, ‖f x‖ ≤ (1 + ‖x‖) ^ (-(n : ℝ)) * - (2 ^ n * ((Finset.Iic m).sup (fun m' => SchwartzMap.seminorm 𝕜 m'.1 m'.2) f)) := by - intro x - rw [rpow_neg (by positivity), ← div_eq_inv_mul, le_div_iff₀' (by positivity), rpow_natCast] - simpa using one_add_le_sup_seminorm_apply (m := m) (k := n) (n := 0) le_rfl le_rfl f x - apply (integral_mono (by simpa using f.integrable_pow_mul μ 0) _ h').trans - · rw [integral_mul_right, ← mul_assoc, mul_comm (2 ^ n)] - rfl - apply h.mul_const) +def integralCLM : 𝓢(D, V) →L[𝕜] V := by + refine mkCLMtoNormedSpace (∫ x, · x ∂μ) + (fun f g ↦ integral_add f.integrable g.integrable) (integral_smul · ·) ?_ + rcases hμ.exists_integrable with ⟨n, h⟩ + let m := (n, 0) + use Finset.Iic m, 2 ^ n * ∫ x : D, (1 + ‖x‖) ^ (- (n : ℝ)) ∂μ + refine ⟨by positivity, fun f ↦ (norm_integral_le_integral_norm f).trans ?_⟩ + have h' : ∀ x, ‖f x‖ ≤ (1 + ‖x‖) ^ (-(n : ℝ)) * + (2 ^ n * ((Finset.Iic m).sup (fun m' => SchwartzMap.seminorm 𝕜 m'.1 m'.2) f)) := by + intro x + rw [rpow_neg (by positivity), ← div_eq_inv_mul, le_div_iff₀' (by positivity), rpow_natCast] + simpa using one_add_le_sup_seminorm_apply (m := m) (k := n) (n := 0) le_rfl le_rfl f x + apply (integral_mono (by simpa using f.integrable_pow_mul μ 0) _ h').trans + · rw [integral_mul_right, ← mul_assoc, mul_comm (2 ^ n)] + rfl + apply h.mul_const variable (𝕜) in @[simp] From 2bb0410843932d03ff2c12cd9bf2e996f514a36b Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Fri, 22 Nov 2024 05:55:44 +0000 Subject: [PATCH 193/829] chore(discover-lean-pr-testing): fix multiline output to GITHUB_OUTPUT (#19355) --- .github/workflows/discover-lean-pr-testing.yml | 4 +--- 1 file changed, 1 insertion(+), 3 deletions(-) diff --git a/.github/workflows/discover-lean-pr-testing.yml b/.github/workflows/discover-lean-pr-testing.yml index 69e46265bbbc0..2bee99c8e421d 100644 --- a/.github/workflows/discover-lean-pr-testing.yml +++ b/.github/workflows/discover-lean-pr-testing.yml @@ -4,8 +4,6 @@ on: push: branches: - nightly-testing - paths: - - lean-toolchain jobs: discover-lean-pr-testing: @@ -69,7 +67,7 @@ jobs: # Output the PRs echo "$PRS" - echo "prs=$PRS" >> "$GITHUB_OUTPUT" + printf "prs<> "$GITHUB_OUTPUT" - name: Use merged PRs information id: find-branches From af1911940422ded96480cde9daed0a5de7a435d3 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Fri, 22 Nov 2024 08:14:07 +0000 Subject: [PATCH 194/829] chore: add shortcut instances after LinearOrder Nat (#19346) This prevents Lean from going through the Lattice structure, which requires additional imports. --- Mathlib/Algebra/Prime/Lemmas.lean | 2 +- Mathlib/CategoryTheory/Functor/OfSequence.lean | 1 - Mathlib/Data/Int/GCD.lean | 2 +- Mathlib/Data/Nat/Defs.lean | 7 +++++++ Mathlib/Data/Nat/Log.lean | 6 +++--- Mathlib/Data/Nat/Prime/Defs.lean | 2 +- Mathlib/Data/Nat/Size.lean | 1 - 7 files changed, 13 insertions(+), 8 deletions(-) diff --git a/Mathlib/Algebra/Prime/Lemmas.lean b/Mathlib/Algebra/Prime/Lemmas.lean index 96ee461f3831f..0211c3bbc8c5e 100644 --- a/Mathlib/Algebra/Prime/Lemmas.lean +++ b/Mathlib/Algebra/Prime/Lemmas.lean @@ -8,7 +8,7 @@ import Mathlib.Algebra.Group.Even import Mathlib.Algebra.Group.Units.Equiv import Mathlib.Algebra.GroupWithZero.Hom import Mathlib.Algebra.Prime.Defs -import Mathlib.Order.Lattice +import Mathlib.Order.Monotone.Basic /-! # Associated, prime, and irreducible elements. diff --git a/Mathlib/CategoryTheory/Functor/OfSequence.lean b/Mathlib/CategoryTheory/Functor/OfSequence.lean index ba70972cf1661..c1f5788cec33b 100644 --- a/Mathlib/CategoryTheory/Functor/OfSequence.lean +++ b/Mathlib/CategoryTheory/Functor/OfSequence.lean @@ -3,7 +3,6 @@ Copyright (c) 2024 Joël Riou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ -import Mathlib.Algebra.Order.Group.Nat import Mathlib.CategoryTheory.Category.Preorder import Mathlib.CategoryTheory.EqToHom diff --git a/Mathlib/Data/Int/GCD.lean b/Mathlib/Data/Int/GCD.lean index eccc02d840c40..e6b7b6ec9f4e8 100644 --- a/Mathlib/Data/Int/GCD.lean +++ b/Mathlib/Data/Int/GCD.lean @@ -7,8 +7,8 @@ import Mathlib.Algebra.Group.Int import Mathlib.Algebra.GroupWithZero.Semiconj import Mathlib.Algebra.Group.Commute.Units import Mathlib.Data.Nat.GCD.Basic -import Mathlib.Order.Lattice import Mathlib.Data.Set.Operations +import Mathlib.Order.Basic import Mathlib.Order.Bounds.Defs /-! diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index 9063d971434ff..2327e2c3bb37a 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -72,6 +72,13 @@ instance instLinearOrder : LinearOrder ℕ where decidableLE := inferInstance decidableEq := inferInstance +-- Shortcut instances +instance : Preorder ℕ := inferInstance +instance : PartialOrder ℕ := inferInstance +instance : Min ℕ := inferInstance +instance : Max ℕ := inferInstance +instance : Ord ℕ := inferInstance + instance instNontrivial : Nontrivial ℕ := ⟨⟨0, 1, Nat.zero_ne_one⟩⟩ @[simp] theorem default_eq_zero : default = 0 := rfl diff --git a/Mathlib/Data/Nat/Log.lean b/Mathlib/Data/Nat/Log.lean index fc11e395a0c41..1d16748dda19e 100644 --- a/Mathlib/Data/Nat/Log.lean +++ b/Mathlib/Data/Nat/Log.lean @@ -3,11 +3,11 @@ Copyright (c) 2020 Simon Hudon. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon, Yaël Dillies -/ +import Mathlib.Order.Interval.Set.Defs +import Mathlib.Order.Monotone.Basic import Mathlib.Tactic.Bound.Attribute -import Mathlib.Tactic.Monotonicity.Attr -import Mathlib.Order.Lattice import Mathlib.Tactic.Contrapose -import Mathlib.Order.Interval.Set.Defs +import Mathlib.Tactic.Monotonicity.Attr /-! # Natural number logarithms diff --git a/Mathlib/Data/Nat/Prime/Defs.lean b/Mathlib/Data/Nat/Prime/Defs.lean index 30b5a96a46888..c978c30544721 100644 --- a/Mathlib/Data/Nat/Prime/Defs.lean +++ b/Mathlib/Data/Nat/Prime/Defs.lean @@ -6,7 +6,7 @@ Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro import Batteries.Data.Nat.Gcd import Mathlib.Algebra.Prime.Defs import Mathlib.Algebra.Ring.Nat -import Mathlib.Order.Lattice +import Mathlib.Order.Basic /-! # Prime numbers diff --git a/Mathlib/Data/Nat/Size.lean b/Mathlib/Data/Nat/Size.lean index fec3892450d06..fd9311b427815 100644 --- a/Mathlib/Data/Nat/Size.lean +++ b/Mathlib/Data/Nat/Size.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Mathlib.Data.Nat.Bits -import Mathlib.Order.Lattice /-! Lemmas about `size`. -/ From 8e3e89e9b3c5b0a719dea497f46946c860381031 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 22 Nov 2024 08:58:22 +0000 Subject: [PATCH 195/829] chore: split Algebra.Group.TypeTags (#19351) --- Mathlib.lean | 4 +- Mathlib/Algebra/Category/Grp/FiniteGrp.lean | 1 + Mathlib/Algebra/FreeMonoid/Count.lean | 1 + Mathlib/Algebra/Group/Action/TypeTags.lean | 2 +- Mathlib/Algebra/Group/Equiv/TypeTags.lean | 2 +- Mathlib/Algebra/Group/Even.lean | 2 +- Mathlib/Algebra/Group/Subgroup/Map.lean | 1 + .../Group/Subsemigroup/Operations.lean | 2 +- .../{TypeTags.lean => TypeTags/Basic.lean} | 124 ++---------------- Mathlib/Algebra/Group/TypeTags/Finite.lean | 28 ++++ Mathlib/Algebra/Group/TypeTags/Hom.lean | 123 +++++++++++++++++ .../Order/Monoid/Unbundled/TypeTags.lean | 2 +- Mathlib/Analysis/Normed/Field/Lemmas.lean | 1 + Mathlib/Data/Finsupp/Defs.lean | 1 + Mathlib/Data/Fintype/Basic.lean | 1 + Mathlib/Data/Nat/Cast/Basic.lean | 1 + Mathlib/Deprecated/Group.lean | 2 +- Mathlib/GroupTheory/FiniteAbelian/Basic.lean | 1 + 18 files changed, 176 insertions(+), 123 deletions(-) rename Mathlib/Algebra/Group/{TypeTags.lean => TypeTags/Basic.lean} (80%) create mode 100644 Mathlib/Algebra/Group/TypeTags/Finite.lean create mode 100644 Mathlib/Algebra/Group/TypeTags/Hom.lean diff --git a/Mathlib.lean b/Mathlib.lean index 5e0aee1763797..55552ad2a0f9b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -321,7 +321,9 @@ import Mathlib.Algebra.Group.Subsemigroup.Membership import Mathlib.Algebra.Group.Subsemigroup.Operations import Mathlib.Algebra.Group.Support import Mathlib.Algebra.Group.Translate -import Mathlib.Algebra.Group.TypeTags +import Mathlib.Algebra.Group.TypeTags.Basic +import Mathlib.Algebra.Group.TypeTags.Finite +import Mathlib.Algebra.Group.TypeTags.Hom import Mathlib.Algebra.Group.ULift import Mathlib.Algebra.Group.UniqueProds.Basic import Mathlib.Algebra.Group.UniqueProds.VectorSpace diff --git a/Mathlib/Algebra/Category/Grp/FiniteGrp.lean b/Mathlib/Algebra/Category/Grp/FiniteGrp.lean index b1f4d7ead61ee..ac01e3ac8ce3c 100644 --- a/Mathlib/Algebra/Category/Grp/FiniteGrp.lean +++ b/Mathlib/Algebra/Category/Grp/FiniteGrp.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Jujian Zhang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang, Nailin Guan, Yuyang Zhao -/ +import Mathlib.Data.Finite.Defs import Mathlib.Algebra.Category.Grp.Basic /-! diff --git a/Mathlib/Algebra/FreeMonoid/Count.lean b/Mathlib/Algebra/FreeMonoid/Count.lean index df4dcfd27f672..53b2c0e4c51c3 100644 --- a/Mathlib/Algebra/FreeMonoid/Count.lean +++ b/Mathlib/Algebra/FreeMonoid/Count.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Algebra.FreeMonoid.Basic +import Mathlib.Algebra.Group.TypeTags.Hom /-! # `List.count` as a bundled homomorphism diff --git a/Mathlib/Algebra/Group/Action/TypeTags.lean b/Mathlib/Algebra/Group/Action/TypeTags.lean index 0fa009b36d114..c6ceefe332560 100644 --- a/Mathlib/Algebra/Group/Action/TypeTags.lean +++ b/Mathlib/Algebra/Group/Action/TypeTags.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Yury Kudryashov -/ import Mathlib.Algebra.Group.Action.End -import Mathlib.Algebra.Group.TypeTags +import Mathlib.Algebra.Group.TypeTags.Hom /-! # Additive and Multiplicative for group actions diff --git a/Mathlib/Algebra/Group/Equiv/TypeTags.lean b/Mathlib/Algebra/Group/Equiv/TypeTags.lean index 78c085f074666..f8e14c0ceb128 100644 --- a/Mathlib/Algebra/Group/Equiv/TypeTags.lean +++ b/Mathlib/Algebra/Group/Equiv/TypeTags.lean @@ -5,7 +5,7 @@ Authors: Johannes Hölzl, Callum Sutton, Yury Kudryashov -/ import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.Algebra.Group.Prod -import Mathlib.Algebra.Group.TypeTags +import Mathlib.Algebra.Group.TypeTags.Hom /-! # Additive and multiplicative equivalences associated to `Multiplicative` and `Additive`. diff --git a/Mathlib/Algebra/Group/Even.lean b/Mathlib/Algebra/Group/Even.lean index 5bd24647c2f65..febd4c5c50c99 100644 --- a/Mathlib/Algebra/Group/Even.lean +++ b/Mathlib/Algebra/Group/Even.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ import Mathlib.Algebra.Group.Opposite -import Mathlib.Algebra.Group.TypeTags +import Mathlib.Algebra.Group.TypeTags.Basic /-! # Squares and even elements diff --git a/Mathlib/Algebra/Group/Subgroup/Map.lean b/Mathlib/Algebra/Group/Subgroup/Map.lean index 84450e745384e..e1219dcd59203 100644 --- a/Mathlib/Algebra/Group/Subgroup/Map.lean +++ b/Mathlib/Algebra/Group/Subgroup/Map.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kexing Ying -/ import Mathlib.Algebra.Group.Subgroup.Lattice +import Mathlib.Algebra.Group.TypeTags.Hom /-! # `map` and `comap` for subgroups diff --git a/Mathlib/Algebra/Group/Subsemigroup/Operations.lean b/Mathlib/Algebra/Group/Subsemigroup/Operations.lean index 8fc45916faae0..2c782614c04cf 100644 --- a/Mathlib/Algebra/Group/Subsemigroup/Operations.lean +++ b/Mathlib/Algebra/Group/Subsemigroup/Operations.lean @@ -6,7 +6,7 @@ Amelia Livingston, Yury Kudryashov, Yakov Pechersky, Jireh Loreaux -/ import Mathlib.Algebra.Group.Prod import Mathlib.Algebra.Group.Subsemigroup.Basic -import Mathlib.Algebra.Group.TypeTags +import Mathlib.Algebra.Group.TypeTags.Basic /-! # Operations on `Subsemigroup`s diff --git a/Mathlib/Algebra/Group/TypeTags.lean b/Mathlib/Algebra/Group/TypeTags/Basic.lean similarity index 80% rename from Mathlib/Algebra/Group/TypeTags.lean rename to Mathlib/Algebra/Group/TypeTags/Basic.lean index f6e2f9c2cc9e0..c506dc1cf3541 100644 --- a/Mathlib/Algebra/Group/TypeTags.lean +++ b/Mathlib/Algebra/Group/TypeTags/Basic.lean @@ -3,8 +3,12 @@ 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.Algebra.Group.Hom.Defs -import Mathlib.Data.Finite.Defs +import Mathlib.Algebra.Group.Pi.Basic +import Mathlib.Data.FunLike.Basic +import Mathlib.Logic.Function.Iterate +import Mathlib.Logic.Equiv.Defs +import Mathlib.Tactic.Set +import Mathlib.Util.AssertExists import Mathlib.Logic.Nontrivial.Basic /-! @@ -32,6 +36,8 @@ This file is similar to `Order.Synonym`. assert_not_exists MonoidWithZero assert_not_exists DenselyOrdered +assert_not_exists MonoidHom +assert_not_exists Finite universe u v @@ -136,16 +142,6 @@ instance [Inhabited α] : Inhabited (Multiplicative α) := instance [Unique α] : Unique (Additive α) := toMul.unique instance [Unique α] : Unique (Multiplicative α) := toAdd.unique -instance [Finite α] : Finite (Additive α) := - Finite.of_equiv α (by rfl) - -instance [Finite α] : Finite (Multiplicative α) := - Finite.of_equiv α (by rfl) - -instance [h : Infinite α] : Infinite (Additive α) := h - -instance [h : Infinite α] : Infinite (Multiplicative α) := h - instance [h : DecidableEq α] : DecidableEq (Multiplicative α) := h instance [h : DecidableEq α] : DecidableEq (Additive α) := h @@ -431,110 +427,6 @@ instance Additive.addCommGroup [CommGroup α] : AddCommGroup (Additive α) := instance Multiplicative.commGroup [AddCommGroup α] : CommGroup (Multiplicative α) := { Multiplicative.group, Multiplicative.commMonoid with } -/-- Reinterpret `α →+ β` as `Multiplicative α →* Multiplicative β`. -/ -@[simps] -def AddMonoidHom.toMultiplicative [AddZeroClass α] [AddZeroClass β] : - (α →+ β) ≃ (Multiplicative α →* Multiplicative β) where - toFun f := { - toFun := fun a => ofAdd (f (toAdd a)) - map_mul' := f.map_add - map_one' := f.map_zero - } - invFun f := { - toFun := fun a => toAdd (f (ofAdd a)) - map_add' := f.map_mul - map_zero' := f.map_one - } - left_inv _ := rfl - right_inv _ := rfl - -@[simp, norm_cast] -lemma AddMonoidHom.coe_toMultiplicative [AddZeroClass α] [AddZeroClass β] (f : α →+ β) : - ⇑(toMultiplicative f) = ofAdd ∘ f ∘ toAdd := rfl - -/-- Reinterpret `α →* β` as `Additive α →+ Additive β`. -/ -@[simps] -def MonoidHom.toAdditive [MulOneClass α] [MulOneClass β] : - (α →* β) ≃ (Additive α →+ Additive β) where - toFun f := { - toFun := fun a => ofMul (f (toMul a)) - map_add' := f.map_mul - map_zero' := f.map_one - } - invFun f := { - toFun := fun a => toMul (f (ofMul a)) - map_mul' := f.map_add - map_one' := f.map_zero - } - left_inv _ := rfl - right_inv _ := rfl - -@[simp, norm_cast] -lemma MonoidHom.coe_toMultiplicative [MulOneClass α] [MulOneClass β] (f : α →* β) : - ⇑(toAdditive f) = ofMul ∘ f ∘ toMul := rfl - -/-- Reinterpret `Additive α →+ β` as `α →* Multiplicative β`. -/ -@[simps] -def AddMonoidHom.toMultiplicative' [MulOneClass α] [AddZeroClass β] : - (Additive α →+ β) ≃ (α →* Multiplicative β) where - toFun f := { - toFun := fun a => ofAdd (f (ofMul a)) - map_mul' := f.map_add - map_one' := f.map_zero - } - invFun f := { - toFun := fun a => toAdd (f (toMul a)) - map_add' := f.map_mul - map_zero' := f.map_one - } - left_inv _ := rfl - right_inv _ := rfl - -@[simp, norm_cast] -lemma AddMonoidHom.coe_toMultiplicative' [MulOneClass α] [AddZeroClass β] (f : Additive α →+ β) : - ⇑(toMultiplicative' f) = ofAdd ∘ f ∘ ofMul := rfl - -/-- Reinterpret `α →* Multiplicative β` as `Additive α →+ β`. -/ -@[simps!] -def MonoidHom.toAdditive' [MulOneClass α] [AddZeroClass β] : - (α →* Multiplicative β) ≃ (Additive α →+ β) := - AddMonoidHom.toMultiplicative'.symm - -@[simp, norm_cast] -lemma MonoidHom.coe_toAdditive' [MulOneClass α] [AddZeroClass β] (f : α →* Multiplicative β) : - ⇑(toAdditive' f) = toAdd ∘ f ∘ toMul := rfl - -/-- Reinterpret `α →+ Additive β` as `Multiplicative α →* β`. -/ -@[simps] -def AddMonoidHom.toMultiplicative'' [AddZeroClass α] [MulOneClass β] : - (α →+ Additive β) ≃ (Multiplicative α →* β) where - toFun f := { - toFun := fun a => toMul (f (toAdd a)) - map_mul' := f.map_add - map_one' := f.map_zero - } - invFun f := { - toFun := fun a => ofMul (f (ofAdd a)) - map_add' := f.map_mul - map_zero' := f.map_one - } - left_inv _ := rfl - right_inv _ := rfl - -@[simp, norm_cast] -lemma AddMonoidHom.coe_toMultiplicative'' [AddZeroClass α] [MulOneClass β] (f : α →+ Additive β) : - ⇑(toMultiplicative'' f) = toMul ∘ f ∘ toAdd := rfl - -/-- Reinterpret `Multiplicative α →* β` as `α →+ Additive β`. -/ -@[simps!] -def MonoidHom.toAdditive'' [AddZeroClass α] [MulOneClass β] : - (Multiplicative α →* β) ≃ (α →+ Additive β) := - AddMonoidHom.toMultiplicative''.symm - -@[simp, norm_cast] -lemma MonoidHom.coe_toAdditive'' [AddZeroClass α] [MulOneClass β] (f : Multiplicative α →* β) : - ⇑(toAdditive'' f) = ofMul ∘ f ∘ ofAdd := rfl - /-- If `α` has some multiplicative structure and coerces to a function, then `Additive α` should also coerce to the same function. diff --git a/Mathlib/Algebra/Group/TypeTags/Finite.lean b/Mathlib/Algebra/Group/TypeTags/Finite.lean new file mode 100644 index 0000000000000..e3885696b373b --- /dev/null +++ b/Mathlib/Algebra/Group/TypeTags/Finite.lean @@ -0,0 +1,28 @@ +/- +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.Algebra.Group.TypeTags.Basic +import Mathlib.Data.Finite.Defs + +/-! +# `Finite` and `Infinite` are preserved by `Additive` and `Multiplicative`. +-/ + +assert_not_exists MonoidWithZero +assert_not_exists DenselyOrdered + +universe u v + +variable {α : Type u} + +instance [Finite α] : Finite (Additive α) := + Finite.of_equiv α (by rfl) + +instance [Finite α] : Finite (Multiplicative α) := + Finite.of_equiv α (by rfl) + +instance [h : Infinite α] : Infinite (Additive α) := h + +instance [h : Infinite α] : Infinite (Multiplicative α) := h diff --git a/Mathlib/Algebra/Group/TypeTags/Hom.lean b/Mathlib/Algebra/Group/TypeTags/Hom.lean new file mode 100644 index 0000000000000..749c54aa7a8b7 --- /dev/null +++ b/Mathlib/Algebra/Group/TypeTags/Hom.lean @@ -0,0 +1,123 @@ +/- +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.Algebra.Group.Hom.Defs +import Mathlib.Algebra.Group.TypeTags.Basic + +/-! +# Transport algebra morphisms between additive and multiplicative types. +-/ + + +universe u v + +variable {α : Type u} {β : Type v} + +open Additive (ofMul toMul) +open Multiplicative (ofAdd toAdd) + +/-- Reinterpret `α →+ β` as `Multiplicative α →* Multiplicative β`. -/ +@[simps] +def AddMonoidHom.toMultiplicative [AddZeroClass α] [AddZeroClass β] : + (α →+ β) ≃ (Multiplicative α →* Multiplicative β) where + toFun f := { + toFun := fun a => ofAdd (f (toAdd a)) + map_mul' := f.map_add + map_one' := f.map_zero + } + invFun f := { + toFun := fun a => toAdd (f (ofAdd a)) + map_add' := f.map_mul + map_zero' := f.map_one + } + left_inv _ := rfl + right_inv _ := rfl + +@[simp, norm_cast] +lemma AddMonoidHom.coe_toMultiplicative [AddZeroClass α] [AddZeroClass β] (f : α →+ β) : + ⇑(toMultiplicative f) = ofAdd ∘ f ∘ toAdd := rfl + +/-- Reinterpret `α →* β` as `Additive α →+ Additive β`. -/ +@[simps] +def MonoidHom.toAdditive [MulOneClass α] [MulOneClass β] : + (α →* β) ≃ (Additive α →+ Additive β) where + toFun f := { + toFun := fun a => ofMul (f (toMul a)) + map_add' := f.map_mul + map_zero' := f.map_one + } + invFun f := { + toFun := fun a => toMul (f (ofMul a)) + map_mul' := f.map_add + map_one' := f.map_zero + } + left_inv _ := rfl + right_inv _ := rfl + +@[simp, norm_cast] +lemma MonoidHom.coe_toMultiplicative [MulOneClass α] [MulOneClass β] (f : α →* β) : + ⇑(toAdditive f) = ofMul ∘ f ∘ toMul := rfl + +/-- Reinterpret `Additive α →+ β` as `α →* Multiplicative β`. -/ +@[simps] +def AddMonoidHom.toMultiplicative' [MulOneClass α] [AddZeroClass β] : + (Additive α →+ β) ≃ (α →* Multiplicative β) where + toFun f := { + toFun := fun a => ofAdd (f (ofMul a)) + map_mul' := f.map_add + map_one' := f.map_zero + } + invFun f := { + toFun := fun a => toAdd (f (toMul a)) + map_add' := f.map_mul + map_zero' := f.map_one + } + left_inv _ := rfl + right_inv _ := rfl + +@[simp, norm_cast] +lemma AddMonoidHom.coe_toMultiplicative' [MulOneClass α] [AddZeroClass β] (f : Additive α →+ β) : + ⇑(toMultiplicative' f) = ofAdd ∘ f ∘ ofMul := rfl + +/-- Reinterpret `α →* Multiplicative β` as `Additive α →+ β`. -/ +@[simps!] +def MonoidHom.toAdditive' [MulOneClass α] [AddZeroClass β] : + (α →* Multiplicative β) ≃ (Additive α →+ β) := + AddMonoidHom.toMultiplicative'.symm + +@[simp, norm_cast] +lemma MonoidHom.coe_toAdditive' [MulOneClass α] [AddZeroClass β] (f : α →* Multiplicative β) : + ⇑(toAdditive' f) = toAdd ∘ f ∘ toMul := rfl + +/-- Reinterpret `α →+ Additive β` as `Multiplicative α →* β`. -/ +@[simps] +def AddMonoidHom.toMultiplicative'' [AddZeroClass α] [MulOneClass β] : + (α →+ Additive β) ≃ (Multiplicative α →* β) where + toFun f := { + toFun := fun a => toMul (f (toAdd a)) + map_mul' := f.map_add + map_one' := f.map_zero + } + invFun f := { + toFun := fun a => ofMul (f (ofAdd a)) + map_add' := f.map_mul + map_zero' := f.map_one + } + left_inv _ := rfl + right_inv _ := rfl + +@[simp, norm_cast] +lemma AddMonoidHom.coe_toMultiplicative'' [AddZeroClass α] [MulOneClass β] (f : α →+ Additive β) : + ⇑(toMultiplicative'' f) = toMul ∘ f ∘ toAdd := rfl + +/-- Reinterpret `Multiplicative α →* β` as `α →+ Additive β`. -/ +@[simps!] +def MonoidHom.toAdditive'' [AddZeroClass α] [MulOneClass β] : + (Multiplicative α →* β) ≃ (α →+ Additive β) := + AddMonoidHom.toMultiplicative''.symm + +@[simp, norm_cast] +lemma MonoidHom.coe_toAdditive'' [AddZeroClass α] [MulOneClass β] (f : Multiplicative α →* β) : + ⇑(toAdditive'' f) = ofMul ∘ f ∘ ofAdd := rfl diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean index 1647d43c6a434..a975f6b4f3384 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean @@ -3,9 +3,9 @@ Copyright (c) 2016 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl -/ -import Mathlib.Algebra.Group.TypeTags import Mathlib.Algebra.Order.Monoid.Unbundled.ExistsOfLE import Mathlib.Order.BoundedOrder +import Mathlib.Algebra.Group.TypeTags.Basic /-! # Ordered monoid structures on `Multiplicative α` and `Additive α`. -/ diff --git a/Mathlib/Analysis/Normed/Field/Lemmas.lean b/Mathlib/Analysis/Normed/Field/Lemmas.lean index e22f72874193c..88affc76dbe77 100644 --- a/Mathlib/Analysis/Normed/Field/Lemmas.lean +++ b/Mathlib/Analysis/Normed/Field/Lemmas.lean @@ -5,6 +5,7 @@ Authors: Patrick Massot, Johannes Hölzl -/ import Mathlib.Algebra.Group.AddChar +import Mathlib.Algebra.Group.TypeTags.Finite import Mathlib.Algebra.Order.Ring.Finset import Mathlib.Analysis.Normed.Field.Basic import Mathlib.Analysis.Normed.Group.Bounded diff --git a/Mathlib/Data/Finsupp/Defs.lean b/Mathlib/Data/Finsupp/Defs.lean index 7ac40e922fd7b..be9894c31daf0 100644 --- a/Mathlib/Data/Finsupp/Defs.lean +++ b/Mathlib/Data/Finsupp/Defs.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.Group.Indicator import Mathlib.Algebra.Group.Submonoid.Basic import Mathlib.Data.Finset.Max import Mathlib.Data.Set.Finite.Basic +import Mathlib.Algebra.Group.TypeTags.Hom /-! # Type of functions with finite support diff --git a/Mathlib/Data/Fintype/Basic.lean b/Mathlib/Data/Fintype/Basic.lean index 3bda18e430ecb..d77529bf9ac3e 100644 --- a/Mathlib/Data/Fintype/Basic.lean +++ b/Mathlib/Data/Fintype/Basic.lean @@ -5,6 +5,7 @@ Authors: Mario Carneiro -/ import Mathlib.Data.Finset.Image import Mathlib.Data.List.FinRange +import Mathlib.Data.Finite.Defs /-! # Finite types diff --git a/Mathlib/Data/Nat/Cast/Basic.lean b/Mathlib/Data/Nat/Cast/Basic.lean index d4f2aa9e4fd08..b8cda137de314 100644 --- a/Mathlib/Data/Nat/Cast/Basic.lean +++ b/Mathlib/Data/Nat/Cast/Basic.lean @@ -6,6 +6,7 @@ Authors: Mario Carneiro import Mathlib.Algebra.Divisibility.Basic import Mathlib.Algebra.Ring.Hom.Defs import Mathlib.Algebra.Ring.Nat +import Mathlib.Algebra.Group.TypeTags.Hom /-! # Cast of natural numbers (additional theorems) diff --git a/Mathlib/Deprecated/Group.lean b/Mathlib/Deprecated/Group.lean index 539782f49e725..66bc94fe1317e 100644 --- a/Mathlib/Deprecated/Group.lean +++ b/Mathlib/Deprecated/Group.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Algebra.Group.Equiv.Basic -import Mathlib.Algebra.Group.TypeTags import Mathlib.Algebra.Group.Units.Hom import Mathlib.Algebra.Ring.Hom.Defs +import Mathlib.Algebra.Group.TypeTags.Basic /-! # Unbundled monoid and group homomorphisms diff --git a/Mathlib/GroupTheory/FiniteAbelian/Basic.lean b/Mathlib/GroupTheory/FiniteAbelian/Basic.lean index def95dc958b3b..38fa544799b25 100644 --- a/Mathlib/GroupTheory/FiniteAbelian/Basic.lean +++ b/Mathlib/GroupTheory/FiniteAbelian/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Pierre-Alexandre Bazin -/ import Mathlib.Algebra.Module.PID +import Mathlib.Algebra.Group.TypeTags.Finite import Mathlib.Data.ZMod.Quotient /-! From 944e3576a3e6deb8a0b61ace4c5df6b80a47435f Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 22 Nov 2024 09:37:25 +0000 Subject: [PATCH 196/829] chore: move def MonoidHom.inverse to earlier defs file (#19348) No import effect, just moving defs to be with other defs. --- Mathlib/Algebra/Group/Equiv/Basic.lean | 31 -------------------------- Mathlib/Algebra/Group/Hom/Defs.lean | 31 ++++++++++++++++++++++++++ MathlibTest/symm.lean | 9 -------- 3 files changed, 31 insertions(+), 40 deletions(-) diff --git a/Mathlib/Algebra/Group/Equiv/Basic.lean b/Mathlib/Algebra/Group/Equiv/Basic.lean index 81aff3106d93e..203ef93e84ead 100644 --- a/Mathlib/Algebra/Group/Equiv/Basic.lean +++ b/Mathlib/Algebra/Group/Equiv/Basic.lean @@ -45,37 +45,6 @@ theorem map_ne_one_iff {f : F} {x : M} : end EmbeddingLike -/-- Makes a `OneHom` inverse from the bijective inverse of a `OneHom` -/ -@[to_additive (attr := simps) - "Make a `ZeroHom` inverse from the bijective inverse of a `ZeroHom`"] -def OneHom.inverse [One M] [One N] - (f : OneHom M N) (g : N → M) - (h₁ : Function.LeftInverse g f) : - OneHom N M := - { toFun := g, - map_one' := by rw [← f.map_one, h₁] } - -/-- Makes a multiplicative inverse from a bijection which preserves multiplication. -/ -@[to_additive (attr := simps) - "Makes an additive inverse from a bijection which preserves addition."] -def MulHom.inverse [Mul M] [Mul N] (f : M →ₙ* N) (g : N → M) - (h₁ : Function.LeftInverse g f) - (h₂ : Function.RightInverse g f) : N →ₙ* M where - toFun := g - map_mul' x y := - calc - g (x * y) = g (f (g x) * f (g y)) := by rw [h₂ x, h₂ y] - _ = g (f (g x * g y)) := by rw [f.map_mul] - _ = g x * g y := h₁ _ - -/-- The inverse of a bijective `MonoidHom` is a `MonoidHom`. -/ -@[to_additive (attr := simps) - "The inverse of a bijective `AddMonoidHom` is an `AddMonoidHom`."] -def MonoidHom.inverse {A B : Type*} [Monoid A] [Monoid B] (f : A →* B) (g : B → A) - (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) : B →* A := - { (f : OneHom A B).inverse g h₁, - (f : A →ₙ* B).inverse g h₁ h₂ with toFun := g } - /-- `AddEquiv α β` is the type of an equiv `α ≃ β` which preserves addition. -/ structure AddEquiv (A B : Type*) [Add A] [Add B] extends A ≃ B, AddHom A B diff --git a/Mathlib/Algebra/Group/Hom/Defs.lean b/Mathlib/Algebra/Group/Hom/Defs.lean index 14f51627f6d9d..f2540ca0f00ab 100644 --- a/Mathlib/Algebra/Group/Hom/Defs.lean +++ b/Mathlib/Algebra/Group/Hom/Defs.lean @@ -843,6 +843,37 @@ protected theorem MonoidHom.map_zpow' [DivInvMonoid M] [DivInvMonoid N] (f : M (hf : ∀ x, f x⁻¹ = (f x)⁻¹) (a : M) (n : ℤ) : f (a ^ n) = f a ^ n := map_zpow' f hf a n +/-- Makes a `OneHom` inverse from the bijective inverse of a `OneHom` -/ +@[to_additive (attr := simps) + "Make a `ZeroHom` inverse from the bijective inverse of a `ZeroHom`"] +def OneHom.inverse [One M] [One N] + (f : OneHom M N) (g : N → M) + (h₁ : Function.LeftInverse g f) : + OneHom N M := + { toFun := g, + map_one' := by rw [← f.map_one, h₁] } + +/-- Makes a multiplicative inverse from a bijection which preserves multiplication. -/ +@[to_additive (attr := simps) + "Makes an additive inverse from a bijection which preserves addition."] +def MulHom.inverse [Mul M] [Mul N] (f : M →ₙ* N) (g : N → M) + (h₁ : Function.LeftInverse g f) + (h₂ : Function.RightInverse g f) : N →ₙ* M where + toFun := g + map_mul' x y := + calc + g (x * y) = g (f (g x) * f (g y)) := by rw [h₂ x, h₂ y] + _ = g (f (g x * g y)) := by rw [f.map_mul] + _ = g x * g y := h₁ _ + +/-- The inverse of a bijective `MonoidHom` is a `MonoidHom`. -/ +@[to_additive (attr := simps) + "The inverse of a bijective `AddMonoidHom` is an `AddMonoidHom`."] +def MonoidHom.inverse {A B : Type*} [Monoid A] [Monoid B] (f : A →* B) (g : B → A) + (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) : B →* A := + { (f : OneHom A B).inverse g h₁, + (f : A →ₙ* B).inverse g h₁ h₂ with toFun := g } + section End namespace Monoid diff --git a/MathlibTest/symm.lean b/MathlibTest/symm.lean index 7d88a90868362..ae72925a5b7a4 100644 --- a/MathlibTest/symm.lean +++ b/MathlibTest/symm.lean @@ -23,15 +23,6 @@ example (a b c : Nat) (ab : a = b) (bc : b = c) : c = a := by symm_saturate apply Eq.trans <;> assumption -def MulHom.inverse [Mul M] [Mul N] (f : M →ₙ* N) (g : N → M) (h₁ : Function.LeftInverse g f) - (h₂ : Function.RightInverse g f) : N →ₙ* M where - toFun := g - map_mul' x y := - calc - g (x * y) = g (f (g x) * f (g y)) := by rw [h₂ x, h₂ y] - _ = g (f (g x * g y)) := by rw [f.map_mul] - _ = g x * g y := h₁ _ - structure MulEquiv (M N : Type u) [Mul M] [Mul N] extends M ≃ N, M →ₙ* N /-- From 1a2513aa6adbca25ceb22c79fea826dd0acb93f3 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 22 Nov 2024 10:20:46 +0000 Subject: [PATCH 197/829] fix: move `Y` notation into a deeper scope (#19292) Otherwise this aggressively claims the `Y` notation in places where the intent may only be to access `Polynomial.X`. An alternative would be to make `Y` an `abbrev`, which would cause it to participate in regular name resolution, allowing local variables to be called `Y`. [Zulip thread](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/Aggressive.20notation.2C.20Y/near/483466695) --- Mathlib/Algebra/Polynomial/Bivariate.lean | 6 ++++-- Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean | 1 + .../EllipticCurve/DivisionPolynomial/Basic.lean | 1 + Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean | 1 + 4 files changed, 7 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/Bivariate.lean b/Mathlib/Algebra/Polynomial/Bivariate.lean index a78e8ffa4e2c8..d9feee54b8b5c 100644 --- a/Mathlib/Algebra/Polynomial/Bivariate.lean +++ b/Mathlib/Algebra/Polynomial/Bivariate.lean @@ -17,10 +17,12 @@ the abbreviation `CC` to view a constant in the base ring `R` as a bivariate pol -/ /-- The notation `Y` for `X` in the `Polynomial` scope. -/ -scoped[Polynomial] notation3:max "Y" => Polynomial.X (R := Polynomial _) +scoped[Polynomial.Bivariate] notation3:max "Y" => Polynomial.X (R := Polynomial _) /-- The notation `R[X][Y]` for `R[X][X]` in the `Polynomial` scope. -/ -scoped[Polynomial] notation3:max R "[X][Y]" => Polynomial (Polynomial R) +scoped[Polynomial.Bivariate] notation3:max R "[X][Y]" => Polynomial (Polynomial R) + +open scoped Polynomial.Bivariate namespace Polynomial diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean index 5a23dde317010..231b6e92006e9 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean @@ -81,6 +81,7 @@ elliptic curve, rational point, affine coordinates -/ open Polynomial +open scoped Polynomial.Bivariate local macro "C_simp" : tactic => `(tactic| simp only [map_ofNat, C_0, C_1, C_neg, C_add, C_sub, C_mul, C_pow]) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean index 5fb7a921a4ae8..15a2fccfcd2b3 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean @@ -95,6 +95,7 @@ elliptic curve, division polynomial, torsion point -/ open Polynomial +open scoped Polynomial.Bivariate local macro "C_simp" : tactic => `(tactic| simp only [map_ofNat, C_0, C_1, C_neg, C_add, C_sub, C_mul, C_pow]) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean index 7020898f0efde..c01a58c13679a 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean @@ -62,6 +62,7 @@ elliptic curve, group law, class group -/ open Ideal nonZeroDivisors Polynomial +open scoped Polynomial.Bivariate local macro "C_simp" : tactic => `(tactic| simp only [map_ofNat, C_0, C_1, C_neg, C_add, C_sub, C_mul, C_pow]) From d511775f5a2ce12b2d54913c5f9ffdc0457bb663 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 22 Nov 2024 10:20:47 +0000 Subject: [PATCH 198/829] docs(RingTheory/Finiteness/Defs): specify namespace of `Finite` (#19359) Currently, `Module.Finite` and `RingHom.Finite` have docstrings that link to `Finite`, which is confusing. --- Mathlib/RingTheory/Finiteness/Defs.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/RingTheory/Finiteness/Defs.lean b/Mathlib/RingTheory/Finiteness/Defs.lean index 289071d1c7064..a56606fa8c8e9 100644 --- a/Mathlib/RingTheory/Finiteness/Defs.lean +++ b/Mathlib/RingTheory/Finiteness/Defs.lean @@ -100,7 +100,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. -/ +/-- A module over a semiring is `Module.Finite` if it is finitely generated as a module. -/ protected class Module.Finite [Semiring R] [AddCommMonoid M] [Module R M] : Prop where out : (⊤ : Submodule R M).FG @@ -142,7 +142,7 @@ namespace RingHom variable {A B C : Type*} [CommRing A] [CommRing B] [CommRing C] -/-- A ring morphism `A →+* B` is `Finite` if `B` is finitely generated as `A`-module. -/ +/-- A ring morphism `A →+* B` is `RingHom.Finite` if `B` is finitely generated as `A`-module. -/ @[algebraize Module.Finite] def Finite (f : A →+* B) : Prop := letI : Algebra A B := f.toAlgebra From dc8fc4ec255093931e2d5370201349d86d3025e2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Martin=20Dvo=C5=99=C3=A1k?= Date: Fri, 22 Nov 2024 10:42:59 +0000 Subject: [PATCH 199/829] feat: `Matrix.abs_det_submatrix_equiv_equiv` (#19332) --- Mathlib/Algebra/Order/Ring/Abs.lean | 4 ++++ Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean | 11 +++++++++++ 2 files changed, 15 insertions(+) diff --git a/Mathlib/Algebra/Order/Ring/Abs.lean b/Mathlib/Algebra/Order/Ring/Abs.lean index 9948d23da4d4b..5552705b81dff 100644 --- a/Mathlib/Algebra/Order/Ring/Abs.lean +++ b/Mathlib/Algebra/Order/Ring/Abs.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.Order.Group.Abs import Mathlib.Algebra.Order.Ring.Basic import Mathlib.Algebra.Order.Ring.Int import Mathlib.Algebra.Ring.Divisibility.Basic +import Mathlib.Algebra.Ring.Int.Units import Mathlib.Data.Nat.Cast.Order.Ring /-! @@ -151,6 +152,9 @@ theorem abs_sub_sq (a b : α) : |a - b| * |a - b| = a * a + b * b - (1 + 1) * a simp only [mul_add, add_comm, add_left_comm, mul_comm, sub_eq_add_neg, mul_one, mul_neg, neg_add_rev, neg_neg, add_assoc] +lemma abs_unit_intCast (a : ℤˣ) : |((a : ℤ) : α)| = 1 := by + cases Int.units_eq_one_or a <;> simp_all + end LinearOrderedCommRing section diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean index 1001a688e6203..2571ccd9ad821 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean @@ -222,6 +222,17 @@ theorem det_submatrix_equiv_self (e : n ≃ m) (A : Matrix m m R) : intro i rw [Equiv.permCongr_apply, Equiv.symm_apply_apply, submatrix_apply] +/-- Permuting rows and columns with two equivalences does not change the absolute value of the +determinant. -/ +@[simp] +theorem abs_det_submatrix_equiv_equiv {R : Type*} [LinearOrderedCommRing R] + (e₁ e₂ : n ≃ m) (A : Matrix m m R) : + |(A.submatrix e₁ e₂).det| = |A.det| := by + have hee : e₂ = e₁.trans (e₁.symm.trans e₂) := by ext; simp + rw [hee] + show |((A.submatrix id (e₁.symm.trans e₂)).submatrix e₁ e₁).det| = |A.det| + rw [Matrix.det_submatrix_equiv_self, Matrix.det_permute', abs_mul, abs_unit_intCast, one_mul] + /-- Reindexing both indices along the same equivalence preserves the determinant. For the `simp` version of this lemma, see `det_submatrix_equiv_self`; this one is unsuitable because From 81381ad8339f8f2b075d0846b3d11fd3ecca4e75 Mon Sep 17 00:00:00 2001 From: "Filippo A. E. Nuccio" Date: Fri, 22 Nov 2024 11:21:14 +0000 Subject: [PATCH 200/829] doc(Algebra/Central/Defs): correct doc (#19363) Correct two typos and a bad syntax in the doc. --- Mathlib/Algebra/Central/Defs.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/Algebra/Central/Defs.lean b/Mathlib/Algebra/Central/Defs.lean index c88a10b8b3aba..363464e22abef 100644 --- a/Mathlib/Algebra/Central/Defs.lean +++ b/Mathlib/Algebra/Central/Defs.lean @@ -19,7 +19,7 @@ is a (not necessarily commutative) `K`-algebra. ## Implementation notes We require the `K`-center of `D` to be smaller than or equal to the smallest subalgebra so that when -we prove something is central, there we don't need to prove `⊥ ≤ center K D` even though this +we prove something is central, we don't need to prove `⊥ ≤ center K D` even though this direction is trivial. ### Central Simple Algebras @@ -34,11 +34,11 @@ but an instance of `[Algebra.IsCentralSimple K D]` would not imply `[IsSimpleRin synthesization orders (`K` cannot be inferred). Thus, to obtain a central simple `K`-algebra `D`, one should use `Algebra.IsCentral K D` and `IsSimpleRing D` separately. -Note that the predicate `Albgera.IsCentral K D` and `IsSimpleRing D` makes sense just for `K` a +Note that the predicate `Algebra.IsCentral K D` and `IsSimpleRing D` makes sense just for `K` a `CommRing` but it doesn't give the right definition for central simple algebra; for a commutative ring base, one should use the theory of Azumaya algebras. In fact ideals of `K` immediately give rise to nontrivial quotients of `D` so there are no central simple algebras in this case according -to our definition, if K is not a field. +to our definition, if `K` is not a field. The theory of central simple algebras really is a theory over fields. Thus to declare a central simple algebra, one should use the following: From 0c44f3226511eec8ec0fbef477cc35921fe0fa2d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Martin=20Dvo=C5=99=C3=A1k?= Date: Fri, 22 Nov 2024 11:41:25 +0000 Subject: [PATCH 201/829] docs(LinearAlgebra/Matrix/Determinant/Basic): det_submatrix_equiv_self (#19365) --- Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean index 2571ccd9ad821..bd291498d65eb 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean @@ -209,7 +209,7 @@ theorem det_permute' (σ : Perm n) (M : Matrix n n R) : (M.submatrix id σ).det = Perm.sign σ * M.det := by rw [← det_transpose, transpose_submatrix, det_permute, det_transpose] -/-- Permuting rows and columns with the same equivalence has no effect. -/ +/-- Permuting rows and columns with the same equivalence does not change the determinant. -/ @[simp] theorem det_submatrix_equiv_self (e : n ≃ m) (A : Matrix m m R) : det (A.submatrix e e) = det A := by From d380d2c97aba4eadfd37855ce5fb64d87caafd3d Mon Sep 17 00:00:00 2001 From: "Filippo A. E. Nuccio" <109365723+Yu-Misaka@users.noreply.github.com> Date: Fri, 22 Nov 2024 13:03:58 +0000 Subject: [PATCH 202/829] feat(FieldTheory.JacobsonNoether) : add proof of the Jacobson-Noether theorem (#16525) In this PR we provide a proof of the Jacobson-Noether theorem following [this blog](https://ysharifi.wordpress.com/2011/09/30/the-jacobson-noether-theorem/) Co-authored by: Filippo A. E. Nuccio @faenuccio Co-authored by: Wanyi He @Blackfeather007 Co-authored by: Huanyu Zheng @Yu-Misaka Co-authored by: Yi Yuan @yuanyi-350 Co-authored by: Weichen Jiao @AlbertJ-314 Co-authored by: Sihan Wu @Old-Turtledove Co-authored by: @FR-vdash-bot Co-authored-by: faenuccio --- Mathlib.lean | 2 + Mathlib/Algebra/Algebra/Basic.lean | 14 +- Mathlib/Algebra/CharP/LinearMaps.lean | 60 +++++++ Mathlib/Algebra/CharP/Subring.lean | 17 +- Mathlib/Algebra/GroupWithZero/Conj.lean | 8 + Mathlib/FieldTheory/JacobsonNoether.lean | 195 +++++++++++++++++++++++ 6 files changed, 290 insertions(+), 6 deletions(-) create mode 100644 Mathlib/Algebra/CharP/LinearMaps.lean create mode 100644 Mathlib/FieldTheory/JacobsonNoether.lean diff --git a/Mathlib.lean b/Mathlib.lean index 55552ad2a0f9b..238444e50113b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -164,6 +164,7 @@ import Mathlib.Algebra.CharP.ExpChar import Mathlib.Algebra.CharP.IntermediateField import Mathlib.Algebra.CharP.Invertible import Mathlib.Algebra.CharP.Lemmas +import Mathlib.Algebra.CharP.LinearMaps import Mathlib.Algebra.CharP.LocalRing import Mathlib.Algebra.CharP.MixedCharZero import Mathlib.Algebra.CharP.Pi @@ -2946,6 +2947,7 @@ import Mathlib.FieldTheory.IsAlgClosed.Classification import Mathlib.FieldTheory.IsAlgClosed.Spectrum import Mathlib.FieldTheory.IsPerfectClosure import Mathlib.FieldTheory.IsSepClosed +import Mathlib.FieldTheory.JacobsonNoether import Mathlib.FieldTheory.KrullTopology import Mathlib.FieldTheory.KummerExtension import Mathlib.FieldTheory.Laurent diff --git a/Mathlib/Algebra/Algebra/Basic.lean b/Mathlib/Algebra/Algebra/Basic.lean index 3e1c923709bb3..98aab66f635a1 100644 --- a/Mathlib/Algebra/Algebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Basic.lean @@ -133,18 +133,24 @@ end CommSemiring section Ring -variable [CommRing R] -variable (R) - /-- A `Semiring` that is an `Algebra` over a commutative ring carries a natural `Ring` structure. See note [reducible non-instances]. -/ -abbrev semiringToRing [Semiring A] [Algebra R A] : Ring A := +abbrev semiringToRing (R : Type*) [CommRing R] [Semiring A] [Algebra R A] : Ring A := { __ := (inferInstance : Semiring A) __ := Module.addCommMonoidToAddCommGroup R intCast := fun z => algebraMap R A z intCast_ofNat := fun z => by simp only [Int.cast_natCast, map_natCast] intCast_negSucc := fun z => by simp } +instance {R : Type*} [Ring R] : Algebra (Subring.center R) R where + toFun := Subtype.val + map_one' := rfl + map_mul' _ _ := rfl + map_zero' := rfl + map_add' _ _ := rfl + commutes' r x := (Subring.mem_center_iff.1 r.2 x).symm + smul_def' _ _ := rfl + end Ring end Algebra diff --git a/Mathlib/Algebra/CharP/LinearMaps.lean b/Mathlib/Algebra/CharP/LinearMaps.lean new file mode 100644 index 0000000000000..36f3baf18a80d --- /dev/null +++ b/Mathlib/Algebra/CharP/LinearMaps.lean @@ -0,0 +1,60 @@ +/- +Copyright (c) 2024 Wanyi He. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Wanyi He, Huanyu Zheng +-/ +import Mathlib.Algebra.CharP.Algebra +/-! +# Characteristic of the ring of linear Maps + +This file contains properties of the characteristic of the ring of linear maps. +The characteristic of the ring of linear maps is determined by its base ring. + +## Main Results + +- `CharP_if` : For a commutative semiring `R` and a `R`-module `M`, + the characteristic of `R` is equal to the characteristic of the `R`-linear + endomorphisms of `M` when `M` contains an element `x` such that + `r • x = 0` implies `r = 0`. + +## Notations + +- `R` is a commutative semiring +- `M` is a `R`-module + +## Implementation Notes + +One can also deduce similar result via `charP_of_injective_ringHom` and + `R → (M →ₗ[R] M) : r ↦ (fun (x : M) ↦ r • x)`. But this will require stronger condition + compared to `CharP_if`. + +-/ + +namespace Module + +variable {R M : Type*} [CommSemiring R] [AddCommMonoid M] [Module R M] + +/-- For a commutative semiring `R` and a `R`-module `M`, if `M` contains an + element `x` such that `r • x = 0` implies `r = 0` (finding such element usually + depends on specific `•`), then the characteristic of `R` is equal to the + characteristic of the `R`-linear endomorphisms of `M`.-/ +theorem charP_end {p : ℕ} [hchar : CharP R p] + (hreduction : ∃ x : M, ∀ r : R, r • x = 0 → r = 0) : CharP (M →ₗ[R] M) p where + cast_eq_zero_iff' n := by + have exact : (n : M →ₗ[R] M) = (n : R) • 1 := by + simp only [Nat.cast_smul_eq_nsmul, nsmul_eq_mul, mul_one] + rw [exact, LinearMap.ext_iff, ← hchar.1] + exact ⟨fun h ↦ Exists.casesOn hreduction fun x hx ↦ hx n (h x), + fun h ↦ (congrArg (fun t ↦ ∀ x, t • x = 0) h).mpr fun x ↦ zero_smul R x⟩ + +end Module + +/-- For a division ring `D` with center `k`, the ring of `k`-linear endomorphisms + of `D` has the same characteristic as `D`-/ +instance {D : Type*} [DivisionRing D] {p : ℕ} [CharP D p] : + CharP (D →ₗ[(Subring.center D)] D) p := + charP_of_injective_ringHom (Algebra.lmul (Subring.center D) D).toRingHom.injective p + +instance {D : Type*} [DivisionRing D] {p : ℕ} [ExpChar D p] : + ExpChar (D →ₗ[Subring.center D] D) p := + expChar_of_injective_ringHom (Algebra.lmul (Subring.center D) D).toRingHom.injective p diff --git a/Mathlib/Algebra/CharP/Subring.lean b/Mathlib/Algebra/CharP/Subring.lean index e6d5faa37044a..23ed075efbeb1 100644 --- a/Mathlib/Algebra/CharP/Subring.lean +++ b/Mathlib/Algebra/CharP/Subring.lean @@ -3,8 +3,7 @@ Copyright (c) 2020 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ -import Mathlib.Algebra.CharP.Defs -import Mathlib.Algebra.Ring.Subring.Defs +import Mathlib.Algebra.CharP.Algebra /-! # Characteristic of subrings @@ -33,4 +32,18 @@ instance subring (R : Type u) [Ring R] (p : ℕ) [CharP R p] (S : Subring R) : C instance subring' (R : Type u) [CommRing R] (p : ℕ) [CharP R p] (S : Subring R) : CharP S p := CharP.subring R p S +/-- The characteristic of a division ring is equal to the characteristic + of its center-/ +theorem charP_center_iff {R : Type u} [Ring R] {p : ℕ} : + CharP (Subring.center R) p ↔ CharP R p := + (algebraMap (Subring.center R) R).charP_iff Subtype.val_injective p + end CharP + +namespace ExpChar + +theorem expChar_center_iff {R : Type u} [Ring R] {p : ℕ} : + ExpChar (Subring.center R) p ↔ ExpChar R p := + (algebraMap (Subring.center R) R).expChar_iff Subtype.val_injective p + +end ExpChar diff --git a/Mathlib/Algebra/GroupWithZero/Conj.lean b/Mathlib/Algebra/GroupWithZero/Conj.lean index 17ccffc430192..6aea4c22972a7 100644 --- a/Mathlib/Algebra/GroupWithZero/Conj.lean +++ b/Mathlib/Algebra/GroupWithZero/Conj.lean @@ -14,9 +14,17 @@ assert_not_exists Multiset -- TODO -- assert_not_exists DenselyOrdered +namespace GroupWithZero + variable {α : Type*} [GroupWithZero α] {a b : α} @[simp] lemma isConj_iff₀ : IsConj a b ↔ ∃ c : α, c ≠ 0 ∧ c * a * c⁻¹ = b := by rw [IsConj, Units.exists_iff_ne_zero (p := (SemiconjBy · a b))] congr! 2 with c exact and_congr_right (mul_inv_eq_iff_eq_mul₀ · |>.symm) + +lemma conj_pow₀ {s : ℕ} {a d : α} (ha : a ≠ 0) : (a⁻¹ * d * a) ^ s = a⁻¹ * d ^ s * a := + let u : αˣ := ⟨a, a⁻¹, mul_inv_cancel₀ ha, inv_mul_cancel₀ ha⟩ + Units.conj_pow' u d s + +end GroupWithZero diff --git a/Mathlib/FieldTheory/JacobsonNoether.lean b/Mathlib/FieldTheory/JacobsonNoether.lean new file mode 100644 index 0000000000000..e019eaf9f4231 --- /dev/null +++ b/Mathlib/FieldTheory/JacobsonNoether.lean @@ -0,0 +1,195 @@ +/- +Copyright (c) 2024 F. Nuccio, H. Zheng, W. He, S. Wu, Y. Yuan, W. Jiao. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Filippo A. E. Nuccio, Huanyu Zheng, Sihan Wu, Wanyi He, Weichen Jiao, Yi Yuan +-/ +import Mathlib.Algebra.Central.Defs +import Mathlib.Algebra.CharP.LinearMaps +import Mathlib.Algebra.CharP.Subring +import Mathlib.Algebra.GroupWithZero.Conj +import Mathlib.Algebra.Lie.OfAssociative +import Mathlib.FieldTheory.PurelyInseparable + +/-! +# The Jacobson-Noether theorem + +This file contains a proof of the Jacobson-Noether theorem and some auxiliary lemmas. +Here we discuss different cases of characteristics of +the noncommutative division algebra `D` with center `k`. + +## Main Results + +- `exists_separable_and_not_isCentral` : (Jacobson-Noether theorem) For a + non-commutative algebraic division algebra `D` (with base ring + being its center `k`), then there exist an element `x` of + `D \ k` that is separable over its center. +- `exists_separable_and_not_isCentral'` : (Jacobson-Noether theorem) For a + non-commutative algebraic division algebra `D` (with base ring + being a field `L`), if the center of `D` over `L` is `L`, + then there exist an element `x` of `D \ L` that is separable over `L`. + +## Notations + +- `D` is a noncommutative division algebra +- `k` is the center of `D` + +## Implementation Notes + +Mathematically, `exists_separable_and_not_isCentral` and `exists_separable_and_not_isCentral'` +are equivalent. + +The difference however, is that the former takes `D` as the only variable +and fixing `D` would forces `k`. Whereas the later takes `D` and `L` as +separate variables constrained by certain relations. + +## Reference +* +-/ + +namespace JacobsonNoether + +variable {D : Type*} [DivisionRing D] [Algebra.IsAlgebraic (Subring.center D) D] + +local notation3 "k" => Subring.center D + +open Polynomial LinearMap LieAlgebra + +/-- If `D` is a purely inseparable extension of `k` with characteristic `p`, + then for every element `a` of `D`, there exists a natural number `n` + such that `a ^ (p ^ n)` is contained in `k`. -/ +lemma exists_pow_mem_center_of_inseparable (p : ℕ) [hchar : ExpChar D p] (a : D) + (hinsep : ∀ x : D, IsSeparable k x → x ∈ k) : ∃ n, a ^ (p ^ n) ∈ k := by + have := (@isPurelyInseparable_iff_pow_mem k D _ _ _ _ p (ExpChar.expChar_center_iff.2 hchar)).1 + have pure : IsPurelyInseparable k D := ⟨Algebra.IsAlgebraic.isIntegral, fun x hx ↦ by + rw [RingHom.mem_range, Subtype.exists] + exact ⟨x, ⟨hinsep x hx, rfl⟩⟩⟩ + obtain ⟨n, ⟨m, hm⟩⟩ := this pure a + have := Subalgebra.range_subset (R := k) ⟨(k).toSubsemiring, fun r ↦ r.2⟩ + exact ⟨n, Set.mem_of_subset_of_mem this <| Set.mem_range.2 ⟨m, hm⟩⟩ + +/-- If `D` is a purely inseparable extension of `k` with characteristic `p`, + then for every element `a` of `D \ k`, there exists a natural number `n` + **greater than 0** such that `a ^ (p ^ n)` is contained in `k`. -/ +lemma exists_pow_mem_center_of_inseparable' (p : ℕ) [ExpChar D p] {a : D} + (ha : a ∉ k) (hinsep : ∀ x : D, IsSeparable k x → x ∈ k) : ∃ n, 1 ≤ n ∧ a ^ (p ^ n) ∈ k := by + obtain ⟨n, hn⟩ := exists_pow_mem_center_of_inseparable p a hinsep + by_cases nzero : n = 0 + · rw [nzero, pow_zero, pow_one] at hn + exact (ha hn).elim + · exact ⟨n, ⟨Nat.one_le_iff_ne_zero.mpr nzero, hn⟩⟩ + +/-- If `D` is a purely inseparable extension of `k` of characteristic `p`, + then for every element `a` of `D \ k`, there exists a natural number `m` + greater than 0 such that `(a * x - x * a) ^ n = 0` (as linear maps) for + every `n` greater than `(p ^ m)`. -/ +lemma exist_pow_eq_zero_of_le (p : ℕ) [hchar : ExpChar D p] + {a : D} (ha : a ∉ k) (hinsep : ∀ x : D, IsSeparable k x → x ∈ k): + ∃ m, 1 ≤ m ∧ ∀ n, p ^ m ≤ n → (ad k D a)^[n] = 0 := by + obtain ⟨m, hm⟩ := exists_pow_mem_center_of_inseparable' p ha hinsep + refine ⟨m, ⟨hm.1, fun n hn ↦ ?_⟩⟩ + have inter : (ad k D a)^[p ^ m] = 0 := by + ext x + rw [ad_eq_lmul_left_sub_lmul_right, ← pow_apply, Pi.sub_apply, + sub_pow_expChar_pow_of_commute p m (commute_mulLeft_right a a), sub_apply, + pow_mulLeft, mulLeft_apply, pow_mulRight, mulRight_apply, Pi.zero_apply, + Subring.mem_center_iff.1 hm.2 x] + exact sub_eq_zero_of_eq rfl + rw [(Nat.sub_eq_iff_eq_add hn).1 rfl, Function.iterate_add, inter, Pi.comp_zero, + iterate_map_zero, Function.const_zero] + +variable (D) in +/-- Jacobson-Noether theorem: For a non-commutative division algebra + `D` that is algebraic over its center `k`, there exists an element + `x` of `D \ k` that is separable over `k`. -/ +theorem exists_separable_and_not_isCentral (H : k ≠ (⊤ : Subring D)) : + ∃ x : D, x ∉ k ∧ IsSeparable k x := by + obtain ⟨p, hp⟩ := ExpChar.exists D + by_contra! insep + replace insep : ∀ x : D, IsSeparable k x → x ∈ k := + fun x h ↦ Classical.byContradiction fun hx ↦ insep x hx h + -- The element `a` below is in `D` but not in `k`. + obtain ⟨a, ha⟩ := not_forall.mp <| mt (Subring.eq_top_iff' k).mpr H + have ha₀ : a ≠ 0 := fun nh ↦ nh ▸ ha <| Subring.zero_mem k + -- We construct another element `b` that does not commute with `a`. + obtain ⟨b, hb1⟩ : ∃ b : D , ad k D a b ≠ 0 := by + rw [Subring.mem_center_iff, not_forall] at ha + use ha.choose + show a * ha.choose - ha.choose * a ≠ 0 + simpa only [ne_eq, sub_eq_zero] using Ne.symm ha.choose_spec + -- We find a maximum natural number `n` such that `(a * x - x * a) ^ n b ≠ 0`. + obtain ⟨n, hn, hb⟩ : ∃ n, 0 < n ∧ (ad k D a)^[n] b ≠ 0 ∧ (ad k D a)^[n + 1] b = 0 := by + obtain ⟨m, -, hm2⟩ := exist_pow_eq_zero_of_le p ha insep + have h_exist : ∃ n, 0 < n ∧ (ad k D a)^[n + 1] b = 0 := ⟨p ^ m, + ⟨expChar_pow_pos D p m, by rw [hm2 (p ^ m + 1) (Nat.le_add_right _ _)]; rfl⟩⟩ + classical + refine ⟨Nat.find h_exist, ⟨(Nat.find_spec h_exist).1, ?_, (Nat.find_spec h_exist).2⟩⟩ + set t := (Nat.find h_exist - 1 : ℕ) with ht + by_cases h_pos : 0 < t + · convert (ne_eq _ _) ▸ not_and.mp (Nat.find_min h_exist (m := t) (by omega)) h_pos + omega + · suffices h_find: Nat.find h_exist = 1 by + rwa [h_find] + rw [not_lt, Nat.le_zero, ht, Nat.sub_eq_zero_iff_le] at h_pos + linarith [(Nat.find_spec h_exist).1] + -- We define `c` to be the value that we proved above to be non-zero. + set c := (ad k D a)^[n] b with hc_def + let _ : Invertible c := ⟨c⁻¹, inv_mul_cancel₀ hb.1, mul_inv_cancel₀ hb.1⟩ + -- We prove that `c` commutes with `a`. + have hc : a * c = c * a := by + apply eq_of_sub_eq_zero + rw [← mulLeft_apply (R := k), ← mulRight_apply (R := k)] + suffices ad k D a c = 0 from by + rw [← this]; rfl + rw [← Function.iterate_succ_apply' (ad k D a) n b, hb.2] + -- We now make some computation to obtain the final equation. + set d := c⁻¹ * a * (ad k D a)^[n - 1] b with hd_def + have hc': c⁻¹ * a = a * c⁻¹ := by + apply_fun (c⁻¹ * · * c⁻¹) at hc + rw [mul_assoc, mul_assoc, mul_inv_cancel₀ hb.1, mul_one, ← mul_assoc, + inv_mul_cancel₀ hb.1, one_mul] at hc + exact hc + have c_eq : a * (ad k D a)^[n - 1] b - (ad k D a)^[n - 1] b * a = c := by + rw [hc_def, ← Nat.sub_add_cancel hn, Function.iterate_succ_apply' (ad k D a) _ b]; rfl + have eq1 : c⁻¹ * a * (ad k D a)^[n - 1] b - c⁻¹ * (ad k D a)^[n - 1] b * a = 1 := by + simp_rw [mul_assoc, (mul_sub_left_distrib c⁻¹ _ _).symm, c_eq, inv_mul_cancel_of_invertible] + -- We show that `a` commutes with `d`. + have deq : a * d - d * a = a := by + nth_rw 3 [← mul_one a] + rw [hd_def, ← eq1, mul_sub, mul_assoc _ _ a, sub_right_inj, hc', + ← mul_assoc, ← mul_assoc, ← mul_assoc] + -- This then yields a contradiction. + apply_fun (a⁻¹ * · ) at deq + rw [mul_sub, ← mul_assoc, inv_mul_cancel₀ ha₀, one_mul, ← mul_assoc, sub_eq_iff_eq_add] at deq + obtain ⟨r, hr⟩ := exists_pow_mem_center_of_inseparable p d insep + apply_fun (· ^ (p ^ r)) at deq + rw [add_pow_expChar_pow_of_commute p r (Commute.one_left _) , one_pow, + GroupWithZero.conj_pow₀ ha₀, ← hr.comm, mul_assoc, inv_mul_cancel₀ ha₀, mul_one, + self_eq_add_left] at deq + exact one_ne_zero deq + +open Subring Algebra in +/-- Jacobson-Noether theorem: For a non-commutative division algebra `D` + that is algebraic over a field `L`, if the center of + `D` coincides with `L`, then there exist an element `x` of `D \ L` + that is separable over `L`. -/ +theorem exists_separable_and_not_isCentral' {L D : Type*} [Field L] [DivisionRing D] + [Algebra L D] [Algebra.IsAlgebraic L D] [Algebra.IsCentral L D] + (hneq : (⊥ : Subalgebra L D) ≠ ⊤) : + ∃ x : D, x ∉ (⊥ : Subalgebra L D) ∧ IsSeparable L x := by + have hcenter : Subalgebra.center L D = ⊥ := le_bot_iff.mp IsCentral.out + have ntrivial : Subring.center D ≠ ⊤ := + congr(Subalgebra.toSubring $hcenter).trans_ne (Subalgebra.toSubring_injective.ne hneq) + set φ := Subalgebra.equivOfEq (⊥ : Subalgebra L D) (.center L D) hcenter.symm + set equiv : L ≃+* (center D) := ((botEquiv L D).symm.trans φ).toRingEquiv + let _ : Algebra L (center D) := equiv.toRingHom.toAlgebra + let _ : Algebra (center D) L := equiv.symm.toRingHom.toAlgebra + have _ : IsScalarTower L (center D) D := .of_algebraMap_eq fun _ ↦ rfl + have _ : IsScalarTower (center D) L D := .of_algebraMap_eq fun x ↦ by + rw [IsScalarTower.algebraMap_apply L (center D)] + congr + exact (equiv.apply_symm_apply x).symm + have _ : Algebra.IsAlgebraic (center D) D := .tower_top (K := L) _ + obtain ⟨x, hxd, hx⟩ := exists_separable_and_not_isCentral D ntrivial + exact ⟨x, ⟨by rwa [← Subalgebra.center_toSubring L, hcenter] at hxd, IsSeparable.tower_top _ hx⟩⟩ + +end JacobsonNoether From be959a9ade921f874584eb5761e7c4f6c82f6651 Mon Sep 17 00:00:00 2001 From: blizzard_inc Date: Fri, 22 Nov 2024 13:33:56 +0000 Subject: [PATCH 203/829] feat(LinearAlgebra/Matrix/Determinant/TotallyUnimodular): iff_fintype (#19366) add a lemma relating the definition of `Matrix.IsTotallyUnimodular` to `Fintype` rather than `Fin`. Co-authored-by: blizzard_inc --- .../Matrix/Determinant/TotallyUnimodular.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean index cd7b6cae09bfd..03da13b5d1757 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean @@ -54,6 +54,18 @@ lemma isTotallyUnimodular_iff (A : Matrix m n R) : A.IsTotallyUnimodular ↔ · intro _ _ _ _ _ apply hA +lemma isTotallyUnimodular_iff_fintype.{w} (A : Matrix m n R) : A.IsTotallyUnimodular ↔ + ∀ (ι : Type w) [Fintype ι] [DecidableEq ι], ∀ f : ι → m, ∀ g : ι → n, + (A.submatrix f g).det ∈ Set.range SignType.cast := by + rw [isTotallyUnimodular_iff] + constructor + · intro hA ι _ _ f g + specialize hA (Fintype.card ι) (f ∘ (Fintype.equivFin ι).symm) (g ∘ (Fintype.equivFin ι).symm) + rwa [←submatrix_submatrix, det_submatrix_equiv_self] at hA + · intro hA k f g + specialize hA (ULift (Fin k)) (f ∘ Equiv.ulift) (g ∘ Equiv.ulift) + rwa [←submatrix_submatrix, det_submatrix_equiv_self] at hA + lemma IsTotallyUnimodular.apply {A : Matrix m n R} (hA : A.IsTotallyUnimodular) (i : m) (j : n) : A i j ∈ Set.range SignType.cast := by From e03117db88529135436883761084450509f14d15 Mon Sep 17 00:00:00 2001 From: damiano Date: Fri, 22 Nov 2024 17:25:12 +0000 Subject: [PATCH 204/829] perf: header linter performs fewer checks (#19260) Trying to recover from the [slowdown](http://speed.lean-fro.org/mathlib4/run-detail/e7b27246-a3e6-496a-b552-ff4b45c7236e/d15f04e158ee396b59b6d37cea21d2ebba0dcf7d) introduced in #18033. --- Mathlib/Tactic/Linter/Header.lean | 21 +++++++++++++++++++-- 1 file changed, 19 insertions(+), 2 deletions(-) diff --git a/Mathlib/Tactic/Linter/Header.lean b/Mathlib/Tactic/Linter/Header.lean index 9f5be314fe2ff..19b353da35598 100644 --- a/Mathlib/Tactic/Linter/Header.lean +++ b/Mathlib/Tactic/Linter/Header.lean @@ -218,6 +218,15 @@ def isInMathlib (modName : Name) : IO Bool := do return (ml.map (·.module == modName)).any (·) else return false +/-- `inMathlibRef` is +* `none` at initialization time; +* `some true` if the `header` linter has already discovered that the current file + is imported in `Mathlib.lean`; +* `some false` if the `header` linter has already discovered that the current file + is *not* imported in `Mathlib.lean`. +-/ +initialize inMathlibRef : IO.Ref (Option Bool) ← IO.mkRef none + /-- The "header" style linter checks that a file starts with ``` @@ -290,9 +299,17 @@ def duplicateImportsCheck (imports : Array Syntax) : CommandElabM Unit := do @[inherit_doc Mathlib.Linter.linter.style.header] def headerLinter : Linter where run := withSetOptionIn fun stx ↦ do let mainModule ← getMainModule + let inMathlib? := ← match ← inMathlibRef.get with + | some d => return d + | none => do + let val ← isInMathlib mainModule + -- We store the answer to the question "is this file in `Mathlib.lean`?" in `inMathlibRef` + -- to avoid recomputing its value on every command. This is a performance optimisation. + inMathlibRef.set (some val) + return val -- The linter skips files not imported in `Mathlib.lean`, to avoid linting "scratch files". - -- However, it is active in the test file `MathlibTest.Header` for the linter itself. - unless (← isInMathlib mainModule) || mainModule == `MathlibTest.Header do return + -- It is however active in the test file `MathlibTest.Header` for the linter itself. + unless inMathlib? || mainModule == `MathlibTest.Header do return unless Linter.getLinterValue linter.style.header (← getOptions) do return if (← get).messages.hasErrors then From 178af37e20ade7b888c1ceb1edbc00dcc331562b Mon Sep 17 00:00:00 2001 From: FMLJohn <417963616@qq.com> Date: Fri, 22 Nov 2024 18:05:10 +0000 Subject: [PATCH 205/829] feat(RingTheory/PowerSeries/Basic): `Polynomial.coe_sub` and `Polynomial.coe_neg` (#19339) These are analogous to the existing `Polynomial.coe_add` and `Polynomial.coe_zero`, about the coercion from `Polynomial` to `PowerSeries`. [Zulip thread](https://leanprover.zulipchat.com/#narrow/channel/217875-Is-there-code-for-X.3F/topic/.60.281.20-.20X.20.3A.20.E2.84.A4.5BX.5D.29.2EtoPowerSeries.20.3D.201.20-.20.28PowerSeries.2EX.20.3A.20.E2.84.A4.E2.9F.A6X.E2.9F.A7.29.60/near/483708229). --- Mathlib/RingTheory/LaurentSeries.lean | 10 +++++----- Mathlib/RingTheory/PowerSeries/Basic.lean | 16 ++++++++++++++++ 2 files changed, 21 insertions(+), 5 deletions(-) diff --git a/Mathlib/RingTheory/LaurentSeries.lean b/Mathlib/RingTheory/LaurentSeries.lean index d543b3fa95654..f8ee85f00a0c0 100644 --- a/Mathlib/RingTheory/LaurentSeries.lean +++ b/Mathlib/RingTheory/LaurentSeries.lean @@ -933,7 +933,7 @@ theorem exists_ratFunc_val_lt (f : K⸨X⸩) (γ : ℤₘ₀ˣ) : · erw [hs, ← F_mul, PowerSeries.coe_pow, PowerSeries.coe_X, RatFunc.coe_mul, zpow_neg, zpow_ofNat, inv_eq_one_div (RatFunc.X ^ s), RatFunc.coe_div, RatFunc.coe_pow, RatFunc.coe_X, RatFunc.coe_one, ← inv_eq_one_div, ← mul_sub, map_mul, map_inv₀, ← PowerSeries.coe_X, - valuation_X_pow, ← hs, ← RatFunc.coe_coe, ← coe_sub, ← coe_algebraMap, + valuation_X_pow, ← hs, ← RatFunc.coe_coe, ← PowerSeries.coe_sub, ← coe_algebraMap, valuation_of_algebraMap, ← Units.val_mk0 (a := ((Multiplicative.ofAdd f.order : Multiplicative ℤ) : ℤₘ₀)), ← hη] apply inv_mul_lt_of_lt_mul₀ @@ -943,8 +943,8 @@ theorem exists_ratFunc_val_lt (f : K⸨X⸩) (γ : ℤₘ₀ˣ) : · obtain ⟨s, hs⟩ := Int.exists_eq_neg_ofNat (Int.neg_nonpos_of_nonneg (not_lt.mp ord_nonpos)) obtain ⟨P, hP⟩ := exists_Polynomial_intValuation_lt (PowerSeries.X ^ s * F) γ use P - erw [← X_order_mul_powerSeriesPart (neg_inj.1 hs).symm, ← RatFunc.coe_coe, ← coe_sub, - ← coe_algebraMap, valuation_of_algebraMap] + erw [← X_order_mul_powerSeriesPart (neg_inj.1 hs).symm, ← RatFunc.coe_coe, + ← PowerSeries.coe_sub, ← coe_algebraMap, valuation_of_algebraMap] exact hP theorem coe_range_dense : DenseRange ((↑) : RatFunc K → K⸨X⸩) := by @@ -981,7 +981,7 @@ theorem inducing_coe : IsUniformInducing ((↑) : RatFunc K → K⸨X⸩) := by refine ⟨⟨d, by rfl⟩, subset_trans (fun _ _ ↦ pre_R ?_) pre_T⟩ apply hd simp only [sub_zero, Set.mem_setOf_eq] - erw [← coe_sub, ← valuation_eq_LaurentSeries_valuation] + erw [← RatFunc.coe_sub, ← valuation_eq_LaurentSeries_valuation] assumption · rintro ⟨_, ⟨hT, pre_T⟩⟩ obtain ⟨d, hd⟩ := Valued.mem_nhds.mp hT @@ -993,7 +993,7 @@ theorem inducing_coe : IsUniformInducing ((↑) : RatFunc K → K⸨X⸩) := by · refine subset_trans (fun _ _ ↦ ?_) pre_T apply hd erw [Set.mem_setOf_eq, sub_zero, valuation_eq_LaurentSeries_valuation, - coe_sub] + RatFunc.coe_sub] assumption theorem continuous_coe : Continuous ((↑) : RatFunc K → K⸨X⸩) := diff --git a/Mathlib/RingTheory/PowerSeries/Basic.lean b/Mathlib/RingTheory/PowerSeries/Basic.lean index 6030481adc2e7..85884f5a24154 100644 --- a/Mathlib/RingTheory/PowerSeries/Basic.lean +++ b/Mathlib/RingTheory/PowerSeries/Basic.lean @@ -758,6 +758,7 @@ namespace Polynomial open Finsupp Polynomial +section CommSemiring variable {R : Type*} [CommSemiring R] (φ ψ : R[X]) -- Porting note: added so we can add the `@[coe]` attribute @@ -877,6 +878,21 @@ theorem coeToPowerSeries.algHom_apply : coeToPowerSeries.algHom A φ = PowerSeries.map (algebraMap R A) ↑φ := rfl +end CommSemiring + +section CommRing +variable {R : Type*} [CommRing R] + +@[simp, norm_cast] +lemma coe_neg (p : R[X]) : ((- p : R[X]) : PowerSeries R) = - p := + coeToPowerSeries.ringHom.map_neg p + +@[simp, norm_cast] +lemma coe_sub (p q : R[X]) : ((p - q : R[X]) : PowerSeries R) = p - q := + coeToPowerSeries.ringHom.map_sub p q + +end CommRing + end Polynomial namespace PowerSeries From 3d401e0bb3b0295605d8083d5683efd6dbbea2eb Mon Sep 17 00:00:00 2001 From: Gio <102917377+gio256@users.noreply.github.com> Date: Fri, 22 Nov 2024 18:29:55 +0000 Subject: [PATCH 206/829] feat(AlgebraicTopology/SimplicialSet): Add auxiliary ext lemma for paths (#19349) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We implement the generalization of the path extensionality lemmas suggested by @joelriou in #19057. Two paths of the same nonzero length can be identified by constructing an identification between each of their arrows. Co-Authored-By: [Joël Riou](https://github.com/joelriou) --- Mathlib/AlgebraicTopology/SimplicialSet/Path.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean index 696d2c7aa9498..047e74c18ff93 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean @@ -73,4 +73,16 @@ lemma spine_map_subinterval {n : ℕ} (j l : ℕ) (hjl : j + l ≤ n) (Δ : X _[ · simp only [spine_arrow, Path.interval, ← FunctorToTypes.map_comp_apply, ← op_comp, mkOfSucc_subinterval_eq] +/-- Two paths of the same nonzero length are equal if all of their arrows are equal. -/ +@[ext] +lemma Path.ext' {n : ℕ} {f g : Path X (n + 1)} + (h : ∀ i : Fin (n + 1), f.arrow i = g.arrow i) : + f = g := by + ext j + · rcases Fin.eq_castSucc_or_eq_last j with ⟨k, hk⟩ | hl + · rw [hk, ← f.arrow_src k, ← g.arrow_src k, h] + · simp only [hl, ← Fin.succ_last] + rw [← f.arrow_tgt (Fin.last n), ← g.arrow_tgt (Fin.last n), h] + · exact h j + end SSet From dd74f0185f0bdda192392376744bfd1a7a9d977b Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Fri, 22 Nov 2024 21:24:50 +0000 Subject: [PATCH 207/829] chore: split Algebra.Group.Nat (#19375) --- Mathlib.lean | 5 +- Mathlib/Algebra/Group/Int.lean | 3 +- Mathlib/Algebra/Group/Nat/Basic.lean | 65 +++++++++++++++ .../Algebra/Group/{Nat.lean => Nat/Even.lean} | 81 +------------------ Mathlib/Algebra/Group/Nat/TypeTags.lean | 24 ++++++ Mathlib/Algebra/Group/Nat/Units.lean | 38 +++++++++ .../Algebra/Group/Submonoid/Operations.lean | 3 +- Mathlib/Algebra/GroupPower/IterateHom.lean | 1 - Mathlib/Algebra/Order/Group/Nat.lean | 2 +- Mathlib/Algebra/Order/Ring/Basic.lean | 1 + Mathlib/Algebra/Ring/Nat.lean | 4 +- Mathlib/Algebra/Ring/Parity.lean | 1 + Mathlib/Data/List/SplitLengths.lean | 2 +- Mathlib/Data/Multiset/Basic.lean | 7 +- Mathlib/Data/Nat/Bits.lean | 3 +- Mathlib/Data/Nat/Bitwise.lean | 2 +- Mathlib/Data/Nat/Cast/Basic.lean | 3 +- Mathlib/Data/Nat/GCD/Basic.lean | 1 + Mathlib/Data/Nat/PSub.lean | 2 +- Mathlib/Data/Nat/Prime/Defs.lean | 2 + Mathlib/Data/Nat/Size.lean | 1 + Mathlib/Data/Set/Enumerate.lean | 2 +- .../SpecificGroups/Quaternion.lean | 1 - Mathlib/Tactic/Sat/FromLRAT.lean | 3 +- MathlibTest/levenshtein.lean | 2 +- 25 files changed, 159 insertions(+), 100 deletions(-) create mode 100644 Mathlib/Algebra/Group/Nat/Basic.lean rename Mathlib/Algebra/Group/{Nat.lean => Nat/Even.lean} (60%) create mode 100644 Mathlib/Algebra/Group/Nat/TypeTags.lean create mode 100644 Mathlib/Algebra/Group/Nat/Units.lean diff --git a/Mathlib.lean b/Mathlib.lean index 238444e50113b..fc3aa1d9991a6 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -281,7 +281,10 @@ import Mathlib.Algebra.Group.Int import Mathlib.Algebra.Group.Invertible.Basic import Mathlib.Algebra.Group.Invertible.Defs import Mathlib.Algebra.Group.MinimalAxioms -import Mathlib.Algebra.Group.Nat +import Mathlib.Algebra.Group.Nat.Basic +import Mathlib.Algebra.Group.Nat.Even +import Mathlib.Algebra.Group.Nat.TypeTags +import Mathlib.Algebra.Group.Nat.Units import Mathlib.Algebra.Group.NatPowAssoc import Mathlib.Algebra.Group.Opposite import Mathlib.Algebra.Group.PNatPowAssoc diff --git a/Mathlib/Algebra/Group/Int.lean b/Mathlib/Algebra/Group/Int.lean index 1fcf3f10ef98c..948531494a03b 100644 --- a/Mathlib/Algebra/Group/Int.lean +++ b/Mathlib/Algebra/Group/Int.lean @@ -3,7 +3,8 @@ Copyright (c) 2016 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad -/ -import Mathlib.Algebra.Group.Nat +import Mathlib.Algebra.Group.Nat.Even +import Mathlib.Algebra.Group.Nat.Units import Mathlib.Algebra.Group.Units.Basic import Mathlib.Data.Int.Sqrt diff --git a/Mathlib/Algebra/Group/Nat/Basic.lean b/Mathlib/Algebra/Group/Nat/Basic.lean new file mode 100644 index 0000000000000..75212737a122b --- /dev/null +++ b/Mathlib/Algebra/Group/Nat/Basic.lean @@ -0,0 +1,65 @@ +/- +Copyright (c) 2014 Floris van Doorn (c) 2016 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro +-/ +import Mathlib.Algebra.Group.Defs + +/-! +# The natural numbers form a monoid + +This file contains the additive and multiplicative monoid instances on the natural numbers. + +See note [foundational algebra order theory]. +-/ + +assert_not_exists MonoidWithZero +assert_not_exists DenselyOrdered + +namespace Nat + +/-! ### Instances -/ + +instance instAddCancelCommMonoid : AddCancelCommMonoid ℕ where + add := Nat.add + add_assoc := Nat.add_assoc + zero := Nat.zero + zero_add := Nat.zero_add + add_zero := Nat.add_zero + add_comm := Nat.add_comm + nsmul m n := m * n + nsmul_zero := Nat.zero_mul + nsmul_succ := succ_mul + add_left_cancel _ _ _ := Nat.add_left_cancel + +instance instCommMonoid : CommMonoid ℕ where + mul := Nat.mul + mul_assoc := Nat.mul_assoc + one := Nat.succ Nat.zero + one_mul := Nat.one_mul + mul_one := Nat.mul_one + mul_comm := Nat.mul_comm + npow m n := n ^ m + npow_zero := Nat.pow_zero + npow_succ _ _ := rfl + +/-! +### Extra instances to short-circuit type class resolution + +These also prevent non-computable instances being used to construct these instances non-computably. +-/ + +instance instAddCommMonoid : AddCommMonoid ℕ := by infer_instance +instance instAddMonoid : AddMonoid ℕ := by infer_instance +instance instMonoid : Monoid ℕ := by infer_instance +instance instCommSemigroup : CommSemigroup ℕ := by infer_instance +instance instSemigroup : Semigroup ℕ := by infer_instance +instance instAddCommSemigroup : AddCommSemigroup ℕ := by infer_instance +instance instAddSemigroup : AddSemigroup ℕ := by infer_instance + +/-! ### Miscellaneous lemmas -/ + +-- We set the simp priority slightly lower than default; later more general lemmas will replace it. +@[simp 900] protected lemma nsmul_eq_mul (m n : ℕ) : m • n = m * n := rfl + +end Nat diff --git a/Mathlib/Algebra/Group/Nat.lean b/Mathlib/Algebra/Group/Nat/Even.lean similarity index 60% rename from Mathlib/Algebra/Group/Nat.lean rename to Mathlib/Algebra/Group/Nat/Even.lean index 42bd89df34d7b..86fcfd3209c0d 100644 --- a/Mathlib/Algebra/Group/Nat.lean +++ b/Mathlib/Algebra/Group/Nat/Even.lean @@ -4,76 +4,18 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Mathlib.Algebra.Group.Even -import Mathlib.Algebra.Group.Units.Defs +import Mathlib.Algebra.Group.Nat.Basic import Mathlib.Data.Nat.Sqrt /-! -# The natural numbers form a monoid - -This file contains the additive and multiplicative monoid instances on the natural numbers. - -See note [foundational algebra order theory]. +# `IsSquare` and `Even` for natural numbers -/ assert_not_exists MonoidWithZero assert_not_exists DenselyOrdered -open Multiplicative - namespace Nat -/-! ### Instances -/ - -instance instAddCancelCommMonoid : AddCancelCommMonoid ℕ where - add := Nat.add - add_assoc := Nat.add_assoc - zero := Nat.zero - zero_add := Nat.zero_add - add_zero := Nat.add_zero - add_comm := Nat.add_comm - nsmul m n := m * n - nsmul_zero := Nat.zero_mul - nsmul_succ := succ_mul - add_left_cancel _ _ _ := Nat.add_left_cancel - -instance instCommMonoid : CommMonoid ℕ where - mul := Nat.mul - mul_assoc := Nat.mul_assoc - one := Nat.succ Nat.zero - one_mul := Nat.one_mul - mul_one := Nat.mul_one - mul_comm := Nat.mul_comm - npow m n := n ^ m - npow_zero := Nat.pow_zero - npow_succ _ _ := rfl - -/-! -### Extra instances to short-circuit type class resolution - -These also prevent non-computable instances being used to construct these instances non-computably. --/ - -instance instAddCommMonoid : AddCommMonoid ℕ := by infer_instance -instance instAddMonoid : AddMonoid ℕ := by infer_instance -instance instMonoid : Monoid ℕ := by infer_instance -instance instCommSemigroup : CommSemigroup ℕ := by infer_instance -instance instSemigroup : Semigroup ℕ := by infer_instance -instance instAddCommSemigroup : AddCommSemigroup ℕ := by infer_instance -instance instAddSemigroup : AddSemigroup ℕ := by infer_instance - -/-! ### Miscellaneous lemmas -/ - --- We set the simp priority slightly lower than default; later more general lemmas will replace it. -@[simp 900] protected lemma nsmul_eq_mul (m n : ℕ) : m • n = m * n := rfl - -section Multiplicative - -lemma toAdd_pow (a : Multiplicative ℕ) (b : ℕ) : toAdd (a ^ b) = toAdd a * b := mul_comm _ _ - -@[simp] lemma ofAdd_mul (a b : ℕ) : ofAdd (a * b) = ofAdd a ^ b := (toAdd_pow _ _).symm - -end Multiplicative - /-! #### Parity -/ variable {m n : ℕ} @@ -158,23 +100,4 @@ example (m n : ℕ) (h : Even m) : ¬Even (n + 3) ↔ Even (m ^ 2 + m + n) := by -- Porting note: the `simp` lemmas about `bit*` no longer apply. example : ¬Even 25394535 := by decide -/-! #### Units -/ - -lemma units_eq_one (u : ℕˣ) : u = 1 := Units.ext <| Nat.eq_one_of_dvd_one ⟨u.inv, u.val_inv.symm⟩ - -lemma addUnits_eq_zero (u : AddUnits ℕ) : u = 0 := - AddUnits.ext <| (Nat.eq_zero_of_add_eq_zero u.val_neg).1 - -@[simp] protected lemma isUnit_iff {n : ℕ} : IsUnit n ↔ n = 1 where - mp := by rintro ⟨u, rfl⟩; obtain rfl := Nat.units_eq_one u; rfl - mpr h := h.symm ▸ ⟨1, rfl⟩ - -instance unique_units : Unique ℕˣ where - default := 1 - uniq := Nat.units_eq_one - -instance unique_addUnits : Unique (AddUnits ℕ) where - default := 0 - uniq := Nat.addUnits_eq_zero - end Nat diff --git a/Mathlib/Algebra/Group/Nat/TypeTags.lean b/Mathlib/Algebra/Group/Nat/TypeTags.lean new file mode 100644 index 0000000000000..0efe1e8dc736b --- /dev/null +++ b/Mathlib/Algebra/Group/Nat/TypeTags.lean @@ -0,0 +1,24 @@ +/- +Copyright (c) 2014 Floris van Doorn (c) 2016 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro +-/ +import Mathlib.Algebra.Group.Nat.Basic +import Mathlib.Algebra.Group.TypeTags.Basic + +/-! +# Lemmas about `Multiplicative ℕ` +-/ + +assert_not_exists MonoidWithZero +assert_not_exists DenselyOrdered + +open Multiplicative + +namespace Nat + +lemma toAdd_pow (a : Multiplicative ℕ) (b : ℕ) : toAdd (a ^ b) = toAdd a * b := mul_comm _ _ + +@[simp] lemma ofAdd_mul (a b : ℕ) : ofAdd (a * b) = ofAdd a ^ b := (toAdd_pow _ _).symm + +end Nat diff --git a/Mathlib/Algebra/Group/Nat/Units.lean b/Mathlib/Algebra/Group/Nat/Units.lean new file mode 100644 index 0000000000000..edc5970d8d2e4 --- /dev/null +++ b/Mathlib/Algebra/Group/Nat/Units.lean @@ -0,0 +1,38 @@ +/- +Copyright (c) 2014 Floris van Doorn (c) 2016 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro +-/ +import Mathlib.Algebra.Group.Nat.Basic +import Mathlib.Algebra.Group.Units.Defs +import Mathlib.Logic.Unique + +/-! +# The unit of the natural numbers +-/ + +assert_not_exists MonoidWithZero +assert_not_exists DenselyOrdered + +namespace Nat + +/-! #### Units -/ + +lemma units_eq_one (u : ℕˣ) : u = 1 := Units.ext <| Nat.eq_one_of_dvd_one ⟨u.inv, u.val_inv.symm⟩ + +lemma addUnits_eq_zero (u : AddUnits ℕ) : u = 0 := + AddUnits.ext <| (Nat.eq_zero_of_add_eq_zero u.val_neg).1 + +@[simp] protected lemma isUnit_iff {n : ℕ} : IsUnit n ↔ n = 1 where + mp := by rintro ⟨u, rfl⟩; obtain rfl := Nat.units_eq_one u; rfl + mpr h := h.symm ▸ ⟨1, rfl⟩ + +instance unique_units : Unique ℕˣ where + default := 1 + uniq := Nat.units_eq_one + +instance unique_addUnits : Unique (AddUnits ℕ) where + default := 0 + uniq := Nat.addUnits_eq_zero + +end Nat diff --git a/Mathlib/Algebra/Group/Submonoid/Operations.lean b/Mathlib/Algebra/Group/Submonoid/Operations.lean index 46c96865817f9..36ee592c01e9d 100644 --- a/Mathlib/Algebra/Group/Submonoid/Operations.lean +++ b/Mathlib/Algebra/Group/Submonoid/Operations.lean @@ -5,9 +5,10 @@ Authors: Johannes Hölzl, Kenny Lau, Johan Commelin, Mario Carneiro, Kevin Buzza Amelia Livingston, Yury Kudryashov -/ import Mathlib.Algebra.Group.Action.Faithful -import Mathlib.Algebra.Group.Nat +import Mathlib.Algebra.Group.Nat.Basic import Mathlib.Algebra.Group.Prod import Mathlib.Algebra.Group.Submonoid.Basic +import Mathlib.Algebra.Group.TypeTags.Basic /-! # Operations on `Submonoid`s diff --git a/Mathlib/Algebra/GroupPower/IterateHom.lean b/Mathlib/Algebra/GroupPower/IterateHom.lean index 1a3170e11ac09..ae097ed297d67 100644 --- a/Mathlib/Algebra/GroupPower/IterateHom.lean +++ b/Mathlib/Algebra/GroupPower/IterateHom.lean @@ -5,7 +5,6 @@ Authors: Yury Kudryashov -/ import Mathlib.Algebra.Group.Action.Opposite import Mathlib.Algebra.Group.Int -import Mathlib.Algebra.Group.Nat import Mathlib.Logic.Function.Iterate import Mathlib.Tactic.Common diff --git a/Mathlib/Algebra/Order/Group/Nat.lean b/Mathlib/Algebra/Order/Group/Nat.lean index 5d7b7628327fd..1385365e53779 100644 --- a/Mathlib/Algebra/Order/Group/Nat.lean +++ b/Mathlib/Algebra/Order/Group/Nat.lean @@ -3,7 +3,7 @@ Copyright (c) 2014 Floris van Doorn (c) 2016 Microsoft Corporation. All rights r Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ -import Mathlib.Algebra.Group.Nat +import Mathlib.Algebra.Group.Nat.Basic import Mathlib.Algebra.Order.Monoid.Canonical.Defs import Mathlib.Algebra.Order.Sub.Defs import Mathlib.Data.Nat.Defs diff --git a/Mathlib/Algebra/Order/Ring/Basic.lean b/Mathlib/Algebra/Order/Ring/Basic.lean index 65844e24d7f30..6f6a50a96337c 100644 --- a/Mathlib/Algebra/Order/Ring/Basic.lean +++ b/Mathlib/Algebra/Order/Ring/Basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2015 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Robert Y. Lewis -/ +import Mathlib.Algebra.Group.Nat.Units import Mathlib.Algebra.Order.Monoid.Unbundled.Pow import Mathlib.Algebra.Order.Ring.Defs import Mathlib.Algebra.Ring.Parity diff --git a/Mathlib/Algebra/Ring/Nat.lean b/Mathlib/Algebra/Ring/Nat.lean index 9b456d31dec8a..e2c57fc380685 100644 --- a/Mathlib/Algebra/Ring/Nat.lean +++ b/Mathlib/Algebra/Ring/Nat.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Mathlib.Algebra.CharZero.Defs -import Mathlib.Algebra.Group.Nat +import Mathlib.Algebra.Group.Nat.Basic import Mathlib.Algebra.Ring.Defs /-! @@ -15,8 +15,6 @@ This file contains the commutative semiring instance on the natural numbers. See note [foundational algebra order theory]. -/ -open Multiplicative - namespace Nat /-! ### Instances -/ diff --git a/Mathlib/Algebra/Ring/Parity.lean b/Mathlib/Algebra/Ring/Parity.lean index 018e1506eb3ec..ff8df726aad0a 100644 --- a/Mathlib/Algebra/Ring/Parity.lean +++ b/Mathlib/Algebra/Ring/Parity.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Damiano Testa. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ +import Mathlib.Algebra.Group.Nat.Even import Mathlib.Data.Nat.Cast.Basic import Mathlib.Data.Nat.Cast.Commute import Mathlib.Data.Set.Operations diff --git a/Mathlib/Data/List/SplitLengths.lean b/Mathlib/Data/List/SplitLengths.lean index e41894fe16e0c..ae9a0052291fa 100644 --- a/Mathlib/Data/List/SplitLengths.lean +++ b/Mathlib/Data/List/SplitLengths.lean @@ -3,8 +3,8 @@ Copyright (c) 2024 Daniel Weber. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Daniel Weber -/ +import Mathlib.Algebra.Group.Nat.Basic import Mathlib.Order.MinMax -import Mathlib.Algebra.Group.Nat /-! # Splitting a list to chunks of specified lengths diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index a594866b4536d..49f764b115ba5 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -3,13 +3,14 @@ Copyright (c) 2015 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Mathlib.Algebra.Group.Nat +import Mathlib.Algebra.Group.Hom.Defs +import Mathlib.Algebra.Group.Nat.Basic import Mathlib.Algebra.Order.Sub.Unbundled.Basic +import Mathlib.Data.List.Perm.Basic +import Mathlib.Data.List.Perm.Lattice import Mathlib.Data.List.Perm.Subperm import Mathlib.Data.Set.List import Mathlib.Order.Hom.Basic -import Mathlib.Data.List.Perm.Lattice -import Mathlib.Data.List.Perm.Basic /-! # Multisets diff --git a/Mathlib/Data/Nat/Bits.lean b/Mathlib/Data/Nat/Bits.lean index 7079602da88d9..fcc0bc2b46077 100644 --- a/Mathlib/Data/Nat/Bits.lean +++ b/Mathlib/Data/Nat/Bits.lean @@ -3,8 +3,7 @@ Copyright (c) 2022 Praneeth Kolichala. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Praneeth Kolichala -/ -import Mathlib.Algebra.Group.Basic -import Mathlib.Algebra.Group.Nat +import Mathlib.Algebra.Group.Nat.Basic import Mathlib.Data.Nat.Defs import Mathlib.Data.Nat.BinaryRec import Mathlib.Data.List.Defs diff --git a/Mathlib/Data/Nat/Bitwise.lean b/Mathlib/Data/Nat/Bitwise.lean index 2de92522a8e9a..e3b32b271b950 100644 --- a/Mathlib/Data/Nat/Bitwise.lean +++ b/Mathlib/Data/Nat/Bitwise.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Markus Himmel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Markus Himmel, Alex Keizer -/ -import Mathlib.Algebra.Group.Units.Basic +import Mathlib.Algebra.Group.Nat.Even import Mathlib.Algebra.NeZero import Mathlib.Algebra.Ring.Nat import Mathlib.Data.List.GetD diff --git a/Mathlib/Data/Nat/Cast/Basic.lean b/Mathlib/Data/Nat/Cast/Basic.lean index b8cda137de314..d9fd72598d7bf 100644 --- a/Mathlib/Data/Nat/Cast/Basic.lean +++ b/Mathlib/Data/Nat/Cast/Basic.lean @@ -4,9 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Mathlib.Algebra.Divisibility.Basic +import Mathlib.Algebra.Group.Even +import Mathlib.Algebra.Group.TypeTags.Hom import Mathlib.Algebra.Ring.Hom.Defs import Mathlib.Algebra.Ring.Nat -import Mathlib.Algebra.Group.TypeTags.Hom /-! # Cast of natural numbers (additional theorems) diff --git a/Mathlib/Data/Nat/GCD/Basic.lean b/Mathlib/Data/Nat/GCD/Basic.lean index ff926c64b4005..2adab886eac34 100644 --- a/Mathlib/Data/Nat/GCD/Basic.lean +++ b/Mathlib/Data/Nat/GCD/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura -/ import Batteries.Data.Nat.Gcd +import Mathlib.Algebra.Group.Nat.Units import Mathlib.Algebra.GroupWithZero.Divisibility import Mathlib.Algebra.Ring.Nat diff --git a/Mathlib/Data/Nat/PSub.lean b/Mathlib/Data/Nat/PSub.lean index e4d0b1efeaaae..dce691789a710 100644 --- a/Mathlib/Data/Nat/PSub.lean +++ b/Mathlib/Data/Nat/PSub.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Mathlib.Algebra.Group.Basic -import Mathlib.Algebra.Group.Nat +import Mathlib.Algebra.Group.Nat.Basic /-! # Partial predecessor and partial subtraction on the natural numbers diff --git a/Mathlib/Data/Nat/Prime/Defs.lean b/Mathlib/Data/Nat/Prime/Defs.lean index c978c30544721..d062858f5ac79 100644 --- a/Mathlib/Data/Nat/Prime/Defs.lean +++ b/Mathlib/Data/Nat/Prime/Defs.lean @@ -4,8 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Batteries.Data.Nat.Gcd +import Mathlib.Algebra.Group.Nat.Units import Mathlib.Algebra.Prime.Defs import Mathlib.Algebra.Ring.Nat +import Mathlib.Data.Nat.Sqrt import Mathlib.Order.Basic /-! diff --git a/Mathlib/Data/Nat/Size.lean b/Mathlib/Data/Nat/Size.lean index fd9311b427815..63929ccf4f5d8 100644 --- a/Mathlib/Data/Nat/Size.lean +++ b/Mathlib/Data/Nat/Size.lean @@ -3,6 +3,7 @@ Copyright (c) 2014 Floris van Doorn (c) 2016 Microsoft Corporation. All rights r Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ +import Mathlib.Algebra.Group.Basic import Mathlib.Data.Nat.Bits /-! Lemmas about `size`. -/ diff --git a/Mathlib/Data/Set/Enumerate.lean b/Mathlib/Data/Set/Enumerate.lean index 908974989a9e3..c5858d0c5e375 100644 --- a/Mathlib/Data/Set/Enumerate.lean +++ b/Mathlib/Data/Set/Enumerate.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ import Mathlib.Algebra.Group.Basic -import Mathlib.Algebra.Group.Nat +import Mathlib.Algebra.Group.Nat.Basic import Mathlib.Tactic.Common import Mathlib.Data.Set.Basic diff --git a/Mathlib/GroupTheory/SpecificGroups/Quaternion.lean b/Mathlib/GroupTheory/SpecificGroups/Quaternion.lean index 3ce9f9ecba4d9..318e8a450260d 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Quaternion.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Quaternion.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Julian Kuelshammer -/ import Mathlib.Data.ZMod.Basic -import Mathlib.Algebra.Group.Nat import Mathlib.Tactic.IntervalCases import Mathlib.GroupTheory.SpecificGroups.Dihedral import Mathlib.GroupTheory.SpecificGroups.Cyclic diff --git a/Mathlib/Tactic/Sat/FromLRAT.lean b/Mathlib/Tactic/Sat/FromLRAT.lean index c1962b10c84a8..db7f18cb54a91 100644 --- a/Mathlib/Tactic/Sat/FromLRAT.lean +++ b/Mathlib/Tactic/Sat/FromLRAT.lean @@ -3,7 +3,8 @@ Copyright (c) 2022 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Mathlib.Algebra.Group.Nat +import Mathlib.Algebra.Group.Nat.Basic +import Mathlib.Tactic.ByContra /-! # `lrat_proof` command diff --git a/MathlibTest/levenshtein.lean b/MathlibTest/levenshtein.lean index 5d2c2d34f0d48..53932d20b2d56 100644 --- a/MathlibTest/levenshtein.lean +++ b/MathlibTest/levenshtein.lean @@ -1,5 +1,5 @@ import Mathlib.Data.List.EditDistance.Defs -import Mathlib.Algebra.Group.Nat +import Mathlib.Algebra.Group.Nat.Basic #guard (suffixLevenshtein Levenshtein.defaultCost "kitten".toList "sitting".toList).1 = From 9e7e42735e19836fc0b9bdccdedee5821afdc982 Mon Sep 17 00:00:00 2001 From: blizzard_inc Date: Fri, 22 Nov 2024 22:03:23 +0000 Subject: [PATCH 208/829] chore(Algebra/Order/SuccPred/TypeTags): fix defeq abuse with Additive and Multiplicative (#19379) these simp lemmas now don't use defeq abuse between `A` and `Multiplicative A`. --- Mathlib/Algebra/Order/SuccPred/TypeTags.lean | 12 ++++++++---- 1 file changed, 8 insertions(+), 4 deletions(-) diff --git a/Mathlib/Algebra/Order/SuccPred/TypeTags.lean b/Mathlib/Algebra/Order/SuccPred/TypeTags.lean index 7c69e6dc40162..6cda0bda7a3e5 100644 --- a/Mathlib/Algebra/Order/SuccPred/TypeTags.lean +++ b/Mathlib/Algebra/Order/SuccPred/TypeTags.lean @@ -36,15 +36,19 @@ namespace Order open Additive Multiplicative @[simp] lemma succ_ofMul [Preorder X] [SuccOrder X] (x : X) : succ (ofMul x) = ofMul (succ x) := rfl -@[simp] lemma succ_toMul [Preorder X] [SuccOrder X] (x : X) : succ (toMul x) = toMul (succ x) := rfl +@[simp] lemma succ_toMul [Preorder X] [SuccOrder X] (x : Additive X) : + succ (toMul x) = toMul (succ x) := rfl @[simp] lemma succ_ofAdd [Preorder X] [SuccOrder X] (x : X) : succ (ofAdd x) = ofAdd (succ x) := rfl -@[simp] lemma succ_toAdd [Preorder X] [SuccOrder X] (x : X) : succ (toAdd x) = toAdd (succ x) := rfl +@[simp] lemma succ_toAdd [Preorder X] [SuccOrder X] (x : Multiplicative X) : + succ (toAdd x) = toAdd (succ x) := rfl @[simp] lemma pred_ofMul [Preorder X] [PredOrder X] (x : X) : pred (ofMul x) = ofMul (pred x) := rfl -@[simp] lemma pred_toMul [Preorder X] [PredOrder X] (x : X) : pred (toMul x) = toMul (pred x) := rfl +@[simp] lemma pred_toMul [Preorder X] [PredOrder X] (x : Additive X) : + pred (toMul x) = toMul (pred x) := rfl @[simp] lemma pred_ofAdd [Preorder X] [PredOrder X] (x : X) : pred (ofAdd x) = ofAdd (pred x) := rfl -@[simp] lemma pred_toAdd [Preorder X] [PredOrder X] (x : X) : pred (toAdd x) = toAdd (pred x) := rfl +@[simp] lemma pred_toAdd [Preorder X] [PredOrder X] (x : Multiplicative X) : + pred (toAdd x) = toAdd (pred x) := rfl end Order From 65cc62903990b71383fddf0866f6a21db3826a24 Mon Sep 17 00:00:00 2001 From: grunweg Date: Fri, 22 Nov 2024 22:03:24 +0000 Subject: [PATCH 209/829] chore: give the 'post or update summary comment' job a descriptive name (#19382) It can be useful to have a useful *job* (and not *workflow step*) name, e.g. in case one looks at JSON information in the queueboard data, about failed CI. --- .github/workflows/PR_summary.yml | 1 + 1 file changed, 1 insertion(+) diff --git a/.github/workflows/PR_summary.yml b/.github/workflows/PR_summary.yml index 09e8747fdd764..6b854a59292d0 100644 --- a/.github/workflows/PR_summary.yml +++ b/.github/workflows/PR_summary.yml @@ -5,6 +5,7 @@ on: jobs: build: + name: "post-or-update-summary-comment" runs-on: ubuntu-latest steps: From b0890aa4ad52262811ba72d7f467848c16b5a497 Mon Sep 17 00:00:00 2001 From: grunweg Date: Fri, 22 Nov 2024 22:03:26 +0000 Subject: [PATCH 210/829] chore: rename no_lints_prime_decls.txt to nolints_... (#19383) For consistency with the other nolints files. Suggested in #19275. --- Mathlib/Tactic/Linter/DocPrime.lean | 6 +++--- scripts/README.md | 3 +-- .../{no_lints_prime_decls.txt => nolints_prime_decls.txt} | 0 scripts/technical-debt-metrics.sh | 2 +- 4 files changed, 5 insertions(+), 6 deletions(-) rename scripts/{no_lints_prime_decls.txt => nolints_prime_decls.txt} (100%) diff --git a/Mathlib/Tactic/Linter/DocPrime.lean b/Mathlib/Tactic/Linter/DocPrime.lean index 7221a12e1ad7f..c2a8a38fbcdcf 100644 --- a/Mathlib/Tactic/Linter/DocPrime.lean +++ b/Mathlib/Tactic/Linter/DocPrime.lean @@ -26,7 +26,7 @@ namespace Mathlib.Linter The "docPrime" linter emits a warning on declarations that have no doc-string and whose name ends with a `'`. -The file `scripts/no_lints_prime_decls.txt` contains a list of temporary exceptions to this linter. +The file `scripts/nolints_prime_decls.txt` contains a list of temporary exceptions to this linter. This list should not be appended to, and become emptied over time. -/ register_option linter.docPrime : Bool := { @@ -63,8 +63,8 @@ def docPrimeLinter : Linter where run := withSetOptionIn fun stx ↦ do relative to the unprimed version, or an explanation as to why no better naming scheme \ is possible." if docstring[0][1].getAtomVal.isEmpty && declName.toString.back == '\'' then - if ← System.FilePath.pathExists "scripts/no_lints_prime_decls.txt" then - if (← IO.FS.lines "scripts/no_lints_prime_decls.txt").contains declName.toString then + if ← System.FilePath.pathExists "scripts/nolints_prime_decls.txt" then + if (← IO.FS.lines "scripts/nolints_prime_decls.txt").contains declName.toString then return else Linter.logLint linter.docPrime declId msg diff --git a/scripts/README.md b/scripts/README.md index 8f55ce30822bb..44156380d5e62 100644 --- a/scripts/README.md +++ b/scripts/README.md @@ -80,8 +80,7 @@ to learn about it as well! **Data files with linter exceptions** - `nolints.json` contains exceptions for all `env_linter`s in mathlib. For permanent and deliberate exceptions, add a `@[nolint lintername]` in the .lean file instead. -- `no_lints_prime_decls.txt` - contains temporary exceptions for the `docPrime` linter +- `nolints_prime_decls.txt` contains temporary exceptions for the `docPrime` linter Both of these files should tend to zero over time; please do not add new entries to these files. PRs removing (the need for) entries are welcome. diff --git a/scripts/no_lints_prime_decls.txt b/scripts/nolints_prime_decls.txt similarity index 100% rename from scripts/no_lints_prime_decls.txt rename to scripts/nolints_prime_decls.txt diff --git a/scripts/technical-debt-metrics.sh b/scripts/technical-debt-metrics.sh index eb24ded1ce153..b3f3df14ec9aa 100755 --- a/scripts/technical-debt-metrics.sh +++ b/scripts/technical-debt-metrics.sh @@ -113,7 +113,7 @@ printf '%s|%s\n' "$(grep -c 'docBlame' scripts/nolints.json)" "documentation nol printf '%s|%s\n' "$(git grep '^set_option linter.style.longFile [0-9]*' Mathlib | wc -l)" "large files" printf '%s|%s\n' "$(git grep "^open .*Classical" | grep -v " in$" -c)" "bare open (scoped) Classical" -printf '%s|%s\n' "$(wc -l < scripts/no_lints_prime_decls.txt)" "exceptions for the docPrime linter" +printf '%s|%s\n' "$(wc -l < scripts/nolints_prime_decls.txt)" "exceptions for the docPrime linter" deprecatedFiles="$(git ls-files '**/Deprecated/*.lean' | xargs wc -l | sed 's=^ *==')" From 4fe03b813c934955c13c0e74e9ba02d5e8449d06 Mon Sep 17 00:00:00 2001 From: blizzard_inc Date: Sat, 23 Nov 2024 00:05:11 +0000 Subject: [PATCH 211/829] chore(Algebra/Group/TypeTags): Remove porting notes related to lean4#1910, use new notation everywhere (#19369) part 1: make applied occurances of `Additive.toAdd` and `Multiplicative.toMul` use newly available dot notation. Co-authored-by: Eric Wieser --- Mathlib/Algebra/AddConstMap/Basic.lean | 2 +- Mathlib/Algebra/AddTorsor.lean | 2 +- .../Algebra/BigOperators/Group/Finset.lean | 12 ++-- Mathlib/Algebra/Group/Action/TypeTags.lean | 8 +-- Mathlib/Algebra/Group/AddChar.lean | 6 +- Mathlib/Algebra/Group/Aut.lean | 4 +- Mathlib/Algebra/Group/Equiv/TypeTags.lean | 20 +++--- Mathlib/Algebra/Group/Even.lean | 4 +- Mathlib/Algebra/Group/Int.lean | 4 +- Mathlib/Algebra/Group/Nat/TypeTags.lean | 2 +- .../Algebra/Group/Submonoid/Membership.lean | 2 +- Mathlib/Algebra/Group/TypeTags/Basic.lean | 65 +++++++++---------- Mathlib/Algebra/Group/TypeTags/Hom.lean | 12 ++-- Mathlib/Algebra/MonoidAlgebra/Basic.lean | 2 +- Mathlib/Algebra/MonoidAlgebra/Defs.lean | 4 +- Mathlib/Algebra/MonoidAlgebra/Division.lean | 6 +- Mathlib/Algebra/MonoidAlgebra/Grading.lean | 4 +- Mathlib/Algebra/Order/Archimedean/Basic.lean | 4 +- .../Order/GroupWithZero/Canonical.lean | 8 +-- Mathlib/Algebra/Order/Monoid/ToMulBot.lean | 2 +- .../Order/Monoid/Unbundled/TypeTags.lean | 8 +-- Mathlib/Algebra/Order/SuccPred/TypeTags.lean | 8 +-- .../Analysis/Normed/Group/Constructions.lean | 12 ++-- Mathlib/Data/Complex/Exponential.lean | 4 +- Mathlib/Data/Int/Cast/Lemmas.lean | 6 +- Mathlib/Data/Int/WithZero.lean | 8 +-- Mathlib/Data/Nat/Cast/Basic.lean | 6 +- Mathlib/Data/ZMod/IntUnitsPower.lean | 14 ++-- .../RotationNumber/TranslationNumber.lean | 2 +- .../GroupTheory/FiniteAbelian/Duality.lean | 4 +- .../AffineSpace/AffineEquiv.lean | 2 +- .../NumberField/Units/DirichletTheorem.lean | 4 +- Mathlib/RepresentationTheory/Basic.lean | 2 +- .../GroupCohomology/LowDegree.lean | 2 +- Mathlib/RepresentationTheory/Rep.lean | 2 +- Mathlib/RingTheory/FreeCommRing.lean | 2 +- Mathlib/RingTheory/LaurentSeries.lean | 2 +- Mathlib/Topology/Constructions.lean | 4 +- Mathlib/Topology/EMetricSpace/Defs.lean | 4 +- Mathlib/Topology/MetricSpace/Defs.lean | 8 +-- .../Topology/MetricSpace/IsometricSMul.lean | 12 ++-- 41 files changed, 142 insertions(+), 147 deletions(-) diff --git a/Mathlib/Algebra/AddConstMap/Basic.lean b/Mathlib/Algebra/AddConstMap/Basic.lean index b8a261ee2098e..25685a81e92e9 100644 --- a/Mathlib/Algebra/AddConstMap/Basic.lean +++ b/Mathlib/Algebra/AddConstMap/Basic.lean @@ -433,7 +433,7 @@ variable {G : Type*} [AddMonoid G] {a : G} as a monoid homomorphism from `Multiplicative G` to `G →+c[a, a] G`. -/ @[simps! (config := .asFn)] def addLeftHom : Multiplicative G →* (G →+c[a, a] G) where - toFun c := Multiplicative.toAdd c +ᵥ .id + toFun c := c.toAdd +ᵥ .id map_one' := by ext; apply zero_add map_mul' _ _ := by ext; apply add_assoc diff --git a/Mathlib/Algebra/AddTorsor.lean b/Mathlib/Algebra/AddTorsor.lean index 2ab53a91f84e3..3f6183df46993 100644 --- a/Mathlib/Algebra/AddTorsor.lean +++ b/Mathlib/Algebra/AddTorsor.lean @@ -358,7 +358,7 @@ theorem constVAdd_add (v₁ v₂ : G) : constVAdd P (v₁ + v₂) = constVAdd P /-- `Equiv.constVAdd` as a homomorphism from `Multiplicative G` to `Equiv.perm P` -/ def constVAddHom : Multiplicative G →* Equiv.Perm P where - toFun v := constVAdd P (Multiplicative.toAdd v) + toFun v := constVAdd P (v.toAdd) map_one' := constVAdd_zero G P map_mul' := constVAdd_add P diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index a8d12379bbcce..282177e1ba7c0 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -2197,7 +2197,7 @@ variable [Monoid α] theorem ofMul_list_prod (s : List α) : ofMul s.prod = (s.map ofMul).sum := by simp [ofMul]; rfl @[simp] -theorem toMul_list_sum (s : List (Additive α)) : toMul s.sum = (s.map toMul).prod := by +theorem toMul_list_sum (s : List (Additive α)) : s.sum.toMul = (s.map toMul).prod := by simp [toMul, ofMul]; rfl end Monoid @@ -2210,7 +2210,7 @@ variable [AddMonoid α] theorem ofAdd_list_prod (s : List α) : ofAdd s.sum = (s.map ofAdd).prod := by simp [ofAdd]; rfl @[simp] -theorem toAdd_list_sum (s : List (Multiplicative α)) : toAdd s.prod = (s.map toAdd).sum := by +theorem toAdd_list_sum (s : List (Multiplicative α)) : s.prod.toAdd = (s.map toAdd).sum := by simp [toAdd, ofAdd]; rfl end AddMonoid @@ -2224,7 +2224,7 @@ theorem ofMul_multiset_prod (s : Multiset α) : ofMul s.prod = (s.map ofMul).sum simp [ofMul]; rfl @[simp] -theorem toMul_multiset_sum (s : Multiset (Additive α)) : toMul s.sum = (s.map toMul).prod := by +theorem toMul_multiset_sum (s : Multiset (Additive α)) : s.sum.toMul = (s.map toMul).prod := by simp [toMul, ofMul]; rfl @[simp] @@ -2233,7 +2233,7 @@ theorem ofMul_prod (s : Finset ι) (f : ι → α) : ofMul (∏ i ∈ s, f i) = @[simp] theorem toMul_sum (s : Finset ι) (f : ι → Additive α) : - toMul (∑ i ∈ s, f i) = ∏ i ∈ s, toMul (f i) := + (∑ i ∈ s, f i).toMul = ∏ i ∈ s, (f i).toMul := rfl end CommMonoid @@ -2248,7 +2248,7 @@ theorem ofAdd_multiset_prod (s : Multiset α) : ofAdd s.sum = (s.map ofAdd).prod @[simp] theorem toAdd_multiset_sum (s : Multiset (Multiplicative α)) : - toAdd s.prod = (s.map toAdd).sum := by + s.prod.toAdd = (s.map toAdd).sum := by simp [toAdd, ofAdd]; rfl @[simp] @@ -2257,7 +2257,7 @@ theorem ofAdd_sum (s : Finset ι) (f : ι → α) : ofAdd (∑ i ∈ s, f i) = @[simp] theorem toAdd_prod (s : Finset ι) (f : ι → Multiplicative α) : - toAdd (∏ i ∈ s, f i) = ∑ i ∈ s, toAdd (f i) := + (∏ i ∈ s, f i).toAdd = ∑ i ∈ s, (f i).toAdd := rfl end AddCommMonoid diff --git a/Mathlib/Algebra/Group/Action/TypeTags.lean b/Mathlib/Algebra/Group/Action/TypeTags.lean index c6ceefe332560..1e65ced1a6024 100644 --- a/Mathlib/Algebra/Group/Action/TypeTags.lean +++ b/Mathlib/Algebra/Group/Action/TypeTags.lean @@ -24,15 +24,15 @@ section open Additive Multiplicative -instance Additive.vadd [SMul α β] : VAdd (Additive α) β where vadd a := (toMul a • ·) +instance Additive.vadd [SMul α β] : VAdd (Additive α) β where vadd a := (a.toMul • ·) -instance Multiplicative.smul [VAdd α β] : SMul (Multiplicative α) β where smul a := (toAdd a +ᵥ ·) +instance Multiplicative.smul [VAdd α β] : SMul (Multiplicative α) β where smul a := (a.toAdd +ᵥ ·) -@[simp] lemma toMul_smul [SMul α β] (a) (b : β) : (toMul a : α) • b = a +ᵥ b := rfl +@[simp] lemma toMul_smul [SMul α β] (a:Additive α) (b : β) : (a.toMul : α) • b = a +ᵥ b := rfl @[simp] lemma ofMul_vadd [SMul α β] (a : α) (b : β) : ofMul a +ᵥ b = a • b := rfl -@[simp] lemma toAdd_vadd [VAdd α β] (a) (b : β) : (toAdd a : α) +ᵥ b = a • b := rfl +@[simp] lemma toAdd_vadd [VAdd α β] (a:Multiplicative α) (b : β) : (a.toAdd : α) +ᵥ b = a • b := rfl @[simp] lemma ofAdd_smul [VAdd α β] (a : α) (b : β) : ofAdd a • b = a +ᵥ b := rfl diff --git a/Mathlib/Algebra/Group/AddChar.lean b/Mathlib/Algebra/Group/AddChar.lean index 23505cff4f752..4b0c0f3e8d9bd 100644 --- a/Mathlib/Algebra/Group/AddChar.lean +++ b/Mathlib/Algebra/Group/AddChar.lean @@ -115,7 +115,7 @@ def toMonoidHom (φ : AddChar A M) : Multiplicative A →* M where -- this instance was a bad idea and conflicted with `instFunLike` above @[simp] lemma toMonoidHom_apply (ψ : AddChar A M) (a : Multiplicative A) : - ψ.toMonoidHom a = ψ (Multiplicative.toAdd a) := + ψ.toMonoidHom a = ψ a.toAdd := rfl /-- An additive character maps multiples by natural numbers to powers. -/ @@ -142,7 +142,7 @@ def toMonoidHomEquiv : AddChar A M ≃ (Multiplicative A →* M) where ⇑(toMonoidHomEquiv.symm ψ) = ψ ∘ Multiplicative.ofAdd := rfl @[simp] lemma toMonoidHomEquiv_apply (ψ : AddChar A M) (a : Multiplicative A) : - toMonoidHomEquiv ψ a = ψ (Multiplicative.toAdd a) := rfl + toMonoidHomEquiv ψ a = ψ a.toAdd := rfl @[simp] lemma toMonoidHomEquiv_symm_apply (ψ : Multiplicative A →* M) (a : A) : toMonoidHomEquiv.symm ψ a = ψ (Multiplicative.ofAdd a) := rfl @@ -180,7 +180,7 @@ lemma coe_toAddMonoidHomEquiv (ψ : AddChar A M) : toAddMonoidHomEquiv ψ a = Additive.ofMul (ψ a) := rfl @[simp] lemma toAddMonoidHomEquiv_symm_apply (ψ : A →+ Additive M) (a : A) : - toAddMonoidHomEquiv.symm ψ a = Additive.toMul (ψ a) := rfl + toAddMonoidHomEquiv.symm ψ a = (ψ a).toMul := rfl /-- The trivial additive character (sending everything to `1`). -/ instance instOne : One (AddChar A M) := toMonoidHomEquiv.one diff --git a/Mathlib/Algebra/Group/Aut.lean b/Mathlib/Algebra/Group/Aut.lean index ac678be0258c9..012f362054268 100644 --- a/Mathlib/Algebra/Group/Aut.lean +++ b/Mathlib/Algebra/Group/Aut.lean @@ -260,9 +260,9 @@ theorem conj_symm_apply [AddGroup G] (g h : G) : (conj g).symm h = -g + h + g := -- Porting note: the exact translation of this mathlib3 lemma would be`(-conj g) h = -g + h + g`, -- but this no longer pass the simp_nf linter, as the LHS simplifies by `toMul_neg` to --- `(Additive.toMul (conj g))⁻¹`. +-- `(conj g).toMul⁻¹`. @[simp] -theorem conj_inv_apply [AddGroup G] (g h : G) : (Additive.toMul (conj g))⁻¹ h = -g + h + g := +theorem conj_inv_apply [AddGroup G] (g h : G) : (conj g).toMul⁻¹ h = -g + h + g := rfl /-- Isomorphic additive groups have isomorphic automorphism groups. -/ diff --git a/Mathlib/Algebra/Group/Equiv/TypeTags.lean b/Mathlib/Algebra/Group/Equiv/TypeTags.lean index f8e14c0ceb128..12338d2419e5a 100644 --- a/Mathlib/Algebra/Group/Equiv/TypeTags.lean +++ b/Mathlib/Algebra/Group/Equiv/TypeTags.lean @@ -118,8 +118,8 @@ and multiplicative endomorphisms of `Multiplicative A`. -/ @[simps] def MulEquiv.piMultiplicative (K : ι → Type*) [∀ i, Add (K i)] : Multiplicative (∀ i : ι, K i) ≃* (∀ i : ι, Multiplicative (K i)) where - toFun x := fun i ↦ Multiplicative.ofAdd <| Multiplicative.toAdd x i - invFun x := Multiplicative.ofAdd fun i ↦ Multiplicative.toAdd (x i) + toFun x := fun i ↦ Multiplicative.ofAdd <| x.toAdd i + invFun x := Multiplicative.ofAdd fun i ↦ (x i).toAdd left_inv _ := rfl right_inv _ := rfl map_mul' _ _ := rfl @@ -134,8 +134,8 @@ abbrev MulEquiv.funMultiplicative [Add G] : @[simps] def AddEquiv.piAdditive (K : ι → Type*) [∀ i, Mul (K i)] : Additive (∀ i : ι, K i) ≃+ (∀ i : ι, Additive (K i)) where - toFun x := fun i ↦ Additive.ofMul <| Additive.toMul x i - invFun x := Additive.ofMul fun i ↦ Additive.toMul (x i) + toFun x := fun i ↦ Additive.ofMul <| x.toMul i + invFun x := Additive.ofMul fun i ↦ (x i).toMul left_inv _ := rfl right_inv _ := rfl map_add' _ _ := rfl @@ -164,9 +164,9 @@ def MulEquiv.multiplicativeAdditive [MulOneClass H] : Multiplicative (Additive H @[simps] def MulEquiv.prodMultiplicative [Add G] [Add H] : Multiplicative (G × H) ≃* Multiplicative G × Multiplicative H where - toFun x := (Multiplicative.ofAdd (Multiplicative.toAdd x).1, - Multiplicative.ofAdd (Multiplicative.toAdd x).2) - invFun := fun (x, y) ↦ Multiplicative.ofAdd (Multiplicative.toAdd x, Multiplicative.toAdd y) + toFun x := (Multiplicative.ofAdd x.toAdd.1, + Multiplicative.ofAdd x.toAdd.2) + invFun := fun (x, y) ↦ Multiplicative.ofAdd (x.toAdd, y.toAdd) left_inv _ := rfl right_inv _ := rfl map_mul' _ _ := rfl @@ -175,9 +175,9 @@ def MulEquiv.prodMultiplicative [Add G] [Add H] : @[simps] def AddEquiv.prodAdditive [Mul G] [Mul H] : Additive (G × H) ≃+ Additive G × Additive H where - toFun x := (Additive.ofMul (Additive.toMul x).1, - Additive.ofMul (Additive.toMul x).2) - invFun := fun (x, y) ↦ Additive.ofMul (Additive.toMul x, Additive.toMul y) + toFun x := (Additive.ofMul x.toMul.1, + Additive.ofMul x.toMul.2) + invFun := fun (x, y) ↦ Additive.ofMul (x.toMul, y.toMul) left_inv _ := rfl right_inv _ := rfl map_add' _ _ := rfl diff --git a/Mathlib/Algebra/Group/Even.lean b/Mathlib/Algebra/Group/Even.lean index febd4c5c50c99..8df0d5b3613c9 100644 --- a/Mathlib/Algebra/Group/Even.lean +++ b/Mathlib/Algebra/Group/Even.lean @@ -62,7 +62,7 @@ instance [DecidablePred (IsSquare : α → Prop)] : DecidablePred (IsSquare : α lemma even_ofMul_iff {a : α} : Even (Additive.ofMul a) ↔ IsSquare a := Iff.rfl @[simp] -lemma isSquare_toMul_iff {a : Additive α} : IsSquare (Additive.toMul a) ↔ Even a := Iff.rfl +lemma isSquare_toMul_iff {a : Additive α} : IsSquare (a.toMul) ↔ Even a := Iff.rfl instance Additive.instDecidablePredEven [DecidablePred (IsSquare : α → Prop)] : DecidablePred (Even : Additive α → Prop) := @@ -76,7 +76,7 @@ variable [Add α] @[simp] lemma isSquare_ofAdd_iff {a : α} : IsSquare (Multiplicative.ofAdd a) ↔ Even a := Iff.rfl @[simp] -lemma even_toAdd_iff {a : Multiplicative α} : Even (Multiplicative.toAdd a) ↔ IsSquare a := Iff.rfl +lemma even_toAdd_iff {a : Multiplicative α} : Even a.toAdd ↔ IsSquare a := Iff.rfl instance Multiplicative.instDecidablePredIsSquare [DecidablePred (Even : α → Prop)] : DecidablePred (IsSquare : Multiplicative α → Prop) := diff --git a/Mathlib/Algebra/Group/Int.lean b/Mathlib/Algebra/Group/Int.lean index 948531494a03b..e9b98108c5110 100644 --- a/Mathlib/Algebra/Group/Int.lean +++ b/Mathlib/Algebra/Group/Int.lean @@ -74,9 +74,9 @@ section Multiplicative open Multiplicative -lemma toAdd_pow (a : Multiplicative ℤ) (b : ℕ) : toAdd (a ^ b) = toAdd a * b := mul_comm _ _ +lemma toAdd_pow (a : Multiplicative ℤ) (b : ℕ) : (a ^ b).toAdd = a.toAdd * b := mul_comm _ _ -lemma toAdd_zpow (a : Multiplicative ℤ) (b : ℤ) : toAdd (a ^ b) = toAdd a * b := mul_comm _ _ +lemma toAdd_zpow (a : Multiplicative ℤ) (b : ℤ) : (a ^ b).toAdd = a.toAdd * b := mul_comm _ _ @[simp] lemma ofAdd_mul (a b : ℤ) : ofAdd (a * b) = ofAdd a ^ b := (toAdd_zpow ..).symm diff --git a/Mathlib/Algebra/Group/Nat/TypeTags.lean b/Mathlib/Algebra/Group/Nat/TypeTags.lean index 0efe1e8dc736b..a312f69a4aab7 100644 --- a/Mathlib/Algebra/Group/Nat/TypeTags.lean +++ b/Mathlib/Algebra/Group/Nat/TypeTags.lean @@ -17,7 +17,7 @@ open Multiplicative namespace Nat -lemma toAdd_pow (a : Multiplicative ℕ) (b : ℕ) : toAdd (a ^ b) = toAdd a * b := mul_comm _ _ +lemma toAdd_pow (a : Multiplicative ℕ) (b : ℕ) : (a ^ b).toAdd = a.toAdd * b := mul_comm _ _ @[simp] lemma ofAdd_mul (a b : ℕ) : ofAdd (a * b) = ofAdd a ^ b := (toAdd_pow _ _).symm diff --git a/Mathlib/Algebra/Group/Submonoid/Membership.lean b/Mathlib/Algebra/Group/Submonoid/Membership.lean index cc9d5be5714ae..2f07619bf80ce 100644 --- a/Mathlib/Algebra/Group/Submonoid/Membership.lean +++ b/Mathlib/Algebra/Group/Submonoid/Membership.lean @@ -475,7 +475,7 @@ when it is injective. The inverse is given by the logarithms. -/ @[simps] def powLogEquiv [DecidableEq M] {n : M} (h : Function.Injective fun m : ℕ => n ^ m) : Multiplicative ℕ ≃* powers n where - toFun m := pow n (Multiplicative.toAdd m) + toFun m := pow n m.toAdd invFun m := Multiplicative.ofAdd (log m) left_inv := log_pow_eq_self h right_inv := pow_log_eq_self diff --git a/Mathlib/Algebra/Group/TypeTags/Basic.lean b/Mathlib/Algebra/Group/TypeTags/Basic.lean index c506dc1cf3541..152313ad2ebc6 100644 --- a/Mathlib/Algebra/Group/TypeTags/Basic.lean +++ b/Mathlib/Algebra/Group/TypeTags/Basic.lean @@ -27,11 +27,6 @@ We also define instances `Additive.*` and `Multiplicative.*` that actually trans This file is similar to `Order.Synonym`. -## Porting notes - -- Since bundled morphism applications that rely on `CoeFun` currently don't work, they are ported - as `toFoo a` rather than `a.toFoo` for now. (https://github.com/leanprover/lean4/issues/1910) - -/ assert_not_exists MonoidWithZero @@ -77,7 +72,7 @@ protected lemma «exists» {p : Additive α → Prop} : (∃ a, p a) ↔ ∃ a, /-- Recursion principle for `Additive`, supported by `cases` and `induction`. -/ @[elab_as_elim, cases_eliminator, induction_eliminator] def rec {motive : Additive α → Sort*} (ofMul : ∀ a, motive (ofMul a)) : ∀ a, motive a := - fun a => ofMul (toMul a) + fun a => ofMul (a.toMul) end Additive @@ -107,7 +102,7 @@ protected lemma «exists» {p : Multiplicative α → Prop} : (∃ a, p a) ↔ /-- Recursion principle for `Multiplicative`, supported by `cases` and `induction`. -/ @[elab_as_elim, cases_eliminator, induction_eliminator] def rec {motive : Multiplicative α → Sort*} (ofAdd : ∀ a, motive (ofAdd a)) : ∀ a, motive a := - fun a => ofAdd (toAdd a) + fun a => ofAdd (a.toAdd) end Multiplicative @@ -115,19 +110,19 @@ open Additive (ofMul toMul) open Multiplicative (ofAdd toAdd) @[simp] -theorem toAdd_ofAdd (x : α) : toAdd (ofAdd x) = x := +theorem toAdd_ofAdd (x : α) : (ofAdd x).toAdd = x := rfl @[simp] -theorem ofAdd_toAdd (x : Multiplicative α) : ofAdd (toAdd x) = x := +theorem ofAdd_toAdd (x : Multiplicative α) : ofAdd x.toAdd = x := rfl @[simp] -theorem toMul_ofMul (x : α) : toMul (ofMul x) = x := +theorem toMul_ofMul (x : α) : (ofMul x).toMul = x := rfl @[simp] -theorem ofMul_toMul (x : Additive α) : ofMul (toMul x) = x := +theorem ofMul_toMul (x : Additive α) : ofMul x.toMul = x := rfl instance [Subsingleton α] : Subsingleton (Additive α) := toMul.injective.subsingleton @@ -153,22 +148,22 @@ instance Multiplicative.instNontrivial [Nontrivial α] : Nontrivial (Multiplicat ofAdd.injective.nontrivial instance Additive.add [Mul α] : Add (Additive α) where - add x y := ofMul (toMul x * toMul y) + add x y := ofMul (x.toMul * y.toMul) instance Multiplicative.mul [Add α] : Mul (Multiplicative α) where - mul x y := ofAdd (toAdd x + toAdd y) + mul x y := ofAdd (x.toAdd + y.toAdd) @[simp] theorem ofAdd_add [Add α] (x y : α) : ofAdd (x + y) = ofAdd x * ofAdd y := rfl @[simp] -theorem toAdd_mul [Add α] (x y : Multiplicative α) : toAdd (x * y) = toAdd x + toAdd y := rfl +theorem toAdd_mul [Add α] (x y : Multiplicative α) : (x * y).toAdd = x.toAdd + y.toAdd := rfl @[simp] theorem ofMul_mul [Mul α] (x y : α) : ofMul (x * y) = ofMul x + ofMul y := rfl @[simp] -theorem toMul_add [Mul α] (x y : Additive α) : toMul (x + y) = toMul x * toMul y := rfl +theorem toMul_add [Mul α] (x y : Additive α) : (x + y).toMul = x.toMul * y.toMul := rfl instance Additive.addSemigroup [Semigroup α] : AddSemigroup (Additive α) := { Additive.add with add_assoc := @mul_assoc α _ } @@ -228,11 +223,11 @@ theorem ofMul_one [One α] : @Additive.ofMul α 1 = 0 := rfl theorem ofMul_eq_zero {A : Type*} [One A] {x : A} : Additive.ofMul x = 0 ↔ x = 1 := Iff.rfl @[simp] -theorem toMul_zero [One α] : toMul (0 : Additive α) = 1 := rfl +theorem toMul_zero [One α] : (0 : Additive α).toMul = 1 := rfl @[simp] lemma toMul_eq_one {α : Type*} [One α] {x : Additive α} : - Additive.toMul x = 1 ↔ x = 0 := + x.toMul = 1 ↔ x = 0 := Iff.rfl instance [Zero α] : One (Multiplicative α) := @@ -247,12 +242,12 @@ theorem ofAdd_eq_one {A : Type*} [Zero A] {x : A} : Multiplicative.ofAdd x = 1 Iff.rfl @[simp] -theorem toAdd_one [Zero α] : toAdd (1 : Multiplicative α) = 0 := +theorem toAdd_one [Zero α] : (1 : Multiplicative α).toAdd = 0 := rfl @[simp] lemma toAdd_eq_zero {α : Type*} [Zero α] {x : Multiplicative α} : - Multiplicative.toAdd x = 0 ↔ x = 1 := + x.toAdd = 0 ↔ x = 1 := Iff.rfl instance Additive.addZeroClass [MulOneClass α] : AddZeroClass (Additive α) where @@ -284,7 +279,7 @@ theorem ofMul_pow [Monoid α] (n : ℕ) (a : α) : ofMul (a ^ n) = n • ofMul a rfl @[simp] -theorem toMul_nsmul [Monoid α] (n : ℕ) (a : Additive α) : toMul (n • a) = toMul a ^ n := +theorem toMul_nsmul [Monoid α] (n : ℕ) (a : Additive α) : (n • a).toMul = a.toMul ^ n := rfl @[simp] @@ -292,7 +287,7 @@ theorem ofAdd_nsmul [AddMonoid α] (n : ℕ) (a : α) : ofAdd (n • a) = ofAdd rfl @[simp] -theorem toAdd_pow [AddMonoid α] (a : Multiplicative α) (n : ℕ) : toAdd (a ^ n) = n • toAdd a := +theorem toAdd_pow [AddMonoid α] (a : Multiplicative α) (n : ℕ) : (a ^ n).toAdd = n • a.toAdd := rfl instance Additive.addLeftCancelMonoid [LeftCancelMonoid α] : AddLeftCancelMonoid (Additive α) := @@ -316,39 +311,39 @@ instance Multiplicative.commMonoid [AddCommMonoid α] : CommMonoid (Multiplicati { Multiplicative.monoid, Multiplicative.commSemigroup with } instance Additive.neg [Inv α] : Neg (Additive α) := - ⟨fun x => ofAdd (toMul x)⁻¹⟩ + ⟨fun x => ofAdd x.toMul⁻¹⟩ @[simp] theorem ofMul_inv [Inv α] (x : α) : ofMul x⁻¹ = -ofMul x := rfl @[simp] -theorem toMul_neg [Inv α] (x : Additive α) : toMul (-x) = (toMul x)⁻¹ := +theorem toMul_neg [Inv α] (x : Additive α) : (-x).toMul = x.toMul⁻¹ := rfl instance Multiplicative.inv [Neg α] : Inv (Multiplicative α) := - ⟨fun x => ofMul (-toAdd x)⟩ + ⟨fun x => ofMul (-x.toAdd)⟩ @[simp] theorem ofAdd_neg [Neg α] (x : α) : ofAdd (-x) = (ofAdd x)⁻¹ := rfl @[simp] -theorem toAdd_inv [Neg α] (x : Multiplicative α) : toAdd x⁻¹ = -(toAdd x) := +theorem toAdd_inv [Neg α] (x : Multiplicative α) : x⁻¹.toAdd = -x.toAdd := rfl instance Additive.sub [Div α] : Sub (Additive α) where - sub x y := ofMul (toMul x / toMul y) + sub x y := ofMul (x.toMul / y.toMul) instance Multiplicative.div [Sub α] : Div (Multiplicative α) where - div x y := ofAdd (toAdd x - toAdd y) + div x y := ofAdd (x.toAdd - y.toAdd) @[simp] theorem ofAdd_sub [Sub α] (x y : α) : ofAdd (x - y) = ofAdd x / ofAdd y := rfl @[simp] -theorem toAdd_div [Sub α] (x y : Multiplicative α) : toAdd (x / y) = toAdd x - toAdd y := +theorem toAdd_div [Sub α] (x y : Multiplicative α) : (x / y).toAdd = x.toAdd - y.toAdd := rfl @[simp] @@ -356,7 +351,7 @@ theorem ofMul_div [Div α] (x y : α) : ofMul (x / y) = ofMul x - ofMul y := rfl @[simp] -theorem toMul_sub [Div α] (x y : Additive α) : toMul (x - y) = toMul x / toMul y := +theorem toMul_sub [Div α] (x y : Additive α) : (x - y).toMul = x.toMul / y.toMul := rfl instance Additive.involutiveNeg [InvolutiveInv α] : InvolutiveNeg (Additive α) := @@ -386,7 +381,7 @@ theorem ofMul_zpow [DivInvMonoid α] (z : ℤ) (a : α) : ofMul (a ^ z) = z • rfl @[simp] -theorem toMul_zsmul [DivInvMonoid α] (z : ℤ) (a : Additive α) : toMul (z • a) = toMul a ^ z := +theorem toMul_zsmul [DivInvMonoid α] (z : ℤ) (a : Additive α) : (z • a).toMul = a.toMul ^ z := rfl @[simp] @@ -394,7 +389,7 @@ theorem ofAdd_zsmul [SubNegMonoid α] (z : ℤ) (a : α) : ofAdd (z • a) = ofA rfl @[simp] -theorem toAdd_zpow [SubNegMonoid α] (a : Multiplicative α) (z : ℤ) : toAdd (a ^ z) = z • toAdd a := +theorem toAdd_zpow [SubNegMonoid α] (a : Multiplicative α) (z : ℤ) : (a ^ z).toAdd = z • a.toAdd := rfl instance Additive.subtractionMonoid [DivisionMonoid α] : SubtractionMonoid (Additive α) := @@ -434,8 +429,8 @@ This allows `Additive` to be used on bundled function types with a multiplicativ is often used for composition, without affecting the behavior of the function itself. -/ instance Additive.coeToFun {α : Type*} {β : α → Sort*} [CoeFun α β] : - CoeFun (Additive α) fun a => β (toMul a) := - ⟨fun a => CoeFun.coe (toMul a)⟩ + CoeFun (Additive α) fun a => β a.toMul := + ⟨fun a => CoeFun.coe a.toMul⟩ /-- If `α` has some additive structure and coerces to a function, then `Multiplicative α` should also coerce to the same function. @@ -444,8 +439,8 @@ This allows `Multiplicative` to be used on bundled function types with an additi is often used for composition, without affecting the behavior of the function itself. -/ instance Multiplicative.coeToFun {α : Type*} {β : α → Sort*} [CoeFun α β] : - CoeFun (Multiplicative α) fun a => β (toAdd a) := - ⟨fun a => CoeFun.coe (toAdd a)⟩ + CoeFun (Multiplicative α) fun a => β a.toAdd := + ⟨fun a => CoeFun.coe a.toAdd⟩ lemma Pi.mulSingle_multiplicativeOfAdd_eq {ι : Type*} [DecidableEq ι] {M : ι → Type*} [(i : ι) → AddMonoid (M i)] (i : ι) (a : M i) (j : ι) : diff --git a/Mathlib/Algebra/Group/TypeTags/Hom.lean b/Mathlib/Algebra/Group/TypeTags/Hom.lean index 749c54aa7a8b7..909615b47da55 100644 --- a/Mathlib/Algebra/Group/TypeTags/Hom.lean +++ b/Mathlib/Algebra/Group/TypeTags/Hom.lean @@ -23,12 +23,12 @@ open Multiplicative (ofAdd toAdd) def AddMonoidHom.toMultiplicative [AddZeroClass α] [AddZeroClass β] : (α →+ β) ≃ (Multiplicative α →* Multiplicative β) where toFun f := { - toFun := fun a => ofAdd (f (toAdd a)) + toFun := fun a => ofAdd (f a.toAdd) map_mul' := f.map_add map_one' := f.map_zero } invFun f := { - toFun := fun a => toAdd (f (ofAdd a)) + toFun := fun a => f (ofAdd a) |>.toAdd map_add' := f.map_mul map_zero' := f.map_one } @@ -44,12 +44,12 @@ lemma AddMonoidHom.coe_toMultiplicative [AddZeroClass α] [AddZeroClass β] (f : def MonoidHom.toAdditive [MulOneClass α] [MulOneClass β] : (α →* β) ≃ (Additive α →+ Additive β) where toFun f := { - toFun := fun a => ofMul (f (toMul a)) + toFun := fun a => ofMul (f a.toMul) map_add' := f.map_mul map_zero' := f.map_one } invFun f := { - toFun := fun a => toMul (f (ofMul a)) + toFun := fun a => (f (ofMul a)).toMul map_mul' := f.map_add map_one' := f.map_zero } @@ -70,7 +70,7 @@ def AddMonoidHom.toMultiplicative' [MulOneClass α] [AddZeroClass β] : map_one' := f.map_zero } invFun f := { - toFun := fun a => toAdd (f (toMul a)) + toFun := fun a => (f a.toMul).toAdd map_add' := f.map_mul map_zero' := f.map_one } @@ -96,7 +96,7 @@ lemma MonoidHom.coe_toAdditive' [MulOneClass α] [AddZeroClass β] (f : α →* def AddMonoidHom.toMultiplicative'' [AddZeroClass α] [MulOneClass β] : (α →+ Additive β) ≃ (Multiplicative α →* β) where toFun f := { - toFun := fun a => toMul (f (toAdd a)) + toFun := fun a => (f a.toAdd).toMul map_mul' := f.map_add map_one' := f.map_zero } diff --git a/Mathlib/Algebra/MonoidAlgebra/Basic.lean b/Mathlib/Algebra/MonoidAlgebra/Basic.lean index 07f16b5034f0f..1e0a8e01e33eb 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Basic.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Basic.lean @@ -463,7 +463,7 @@ theorem lift_def (F : Multiplicative G →* A) : @[simp] theorem lift_symm_apply (F : k[G] →ₐ[k] A) (x : Multiplicative G) : - (lift k G A).symm F x = F (single (Multiplicative.toAdd x) 1) := + (lift k G A).symm F x = F (single x.toAdd 1) := rfl theorem lift_of (F : Multiplicative G →* A) (x : Multiplicative G) : diff --git a/Mathlib/Algebra/MonoidAlgebra/Defs.lean b/Mathlib/Algebra/MonoidAlgebra/Defs.lean index 93e53c6c0603b..41e3c6d67e134 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Defs.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Defs.lean @@ -1300,7 +1300,7 @@ end @[simp] theorem of_apply [AddZeroClass G] (a : Multiplicative G) : - of k G a = single (Multiplicative.toAdd a) 1 := + of k G a = single a.toAdd 1 := rfl @[simp] @@ -1324,7 +1324,7 @@ Note the order of the elements of the product are reversed compared to the argum -/ @[simps] def singleHom [AddZeroClass G] : k × Multiplicative G →* k[G] where - toFun a := single (Multiplicative.toAdd a.2) a.1 + toFun a := single a.2.toAdd a.1 map_one' := rfl map_mul' _a _b := single_mul_single.symm diff --git a/Mathlib/Algebra/MonoidAlgebra/Division.lean b/Mathlib/Algebra/MonoidAlgebra/Division.lean index e0f57c1d1a2b2..7b63bd6d1bbd6 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Division.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Division.lean @@ -83,13 +83,13 @@ theorem divOf_add (x : k[G]) (a b : G) : x /ᵒᶠ (a + b) = x /ᵒᶠ a /ᵒᶠ @[simps] noncomputable def divOfHom : Multiplicative G →* AddMonoid.End k[G] where toFun g := - { toFun := fun x => divOf x (Multiplicative.toAdd g) + { toFun := fun x => divOf x g.toAdd map_zero' := zero_divOf _ - map_add' := fun x y => add_divOf x y (Multiplicative.toAdd g) } + map_add' := fun x y => add_divOf x y g.toAdd } map_one' := AddMonoidHom.ext divOf_zero map_mul' g₁ g₂ := AddMonoidHom.ext fun _x => - (congr_arg _ (add_comm (Multiplicative.toAdd g₁) (Multiplicative.toAdd g₂))).trans + (congr_arg _ (add_comm g₁.toAdd g₂.toAdd)).trans (divOf_add _ _ _) theorem of'_mul_divOf (a : G) (x : k[G]) : of' k G a * x /ᵒᶠ a = x := by diff --git a/Mathlib/Algebra/MonoidAlgebra/Grading.lean b/Mathlib/Algebra/MonoidAlgebra/Grading.lean index 39f3e606eae43..66092e7d7d662 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Grading.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Grading.lean @@ -110,8 +110,8 @@ variable [AddMonoid M] [DecidableEq ι] [AddMonoid ι] [CommSemiring R] (f : M def decomposeAux : R[M] →ₐ[R] ⨁ i : ι, gradeBy R f i := AddMonoidAlgebra.lift R M _ { toFun := fun m => - DirectSum.of (fun i : ι => gradeBy R f i) (f (Multiplicative.toAdd m)) - ⟨Finsupp.single (Multiplicative.toAdd m) 1, single_mem_gradeBy _ _ _⟩ + DirectSum.of (fun i : ι => gradeBy R f i) (f m.toAdd) + ⟨Finsupp.single m.toAdd 1, single_mem_gradeBy _ _ _⟩ map_one' := DirectSum.of_eq_of_gradedMonoid_eq (by congr 2 <;> simp) diff --git a/Mathlib/Algebra/Order/Archimedean/Basic.lean b/Mathlib/Algebra/Order/Archimedean/Basic.lean index c2c50a58878f6..bf952ac20407b 100644 --- a/Mathlib/Algebra/Order/Archimedean/Basic.lean +++ b/Mathlib/Algebra/Order/Archimedean/Basic.lean @@ -64,11 +64,11 @@ instance OrderDual.instMulArchimedean [OrderedCommGroup α] [MulArchimedean α] instance Additive.instArchimedean [OrderedCommGroup α] [MulArchimedean α] : Archimedean (Additive α) := - ⟨fun x _ hy ↦ MulArchimedean.arch (toMul x) hy⟩ + ⟨fun x _ hy ↦ MulArchimedean.arch x.toMul hy⟩ instance Multiplicative.instMulArchimedean [OrderedAddCommGroup α] [Archimedean α] : MulArchimedean (Multiplicative α) := - ⟨fun x _ hy ↦ Archimedean.arch (toAdd x) hy⟩ + ⟨fun x _ hy ↦ Archimedean.arch x.toAdd hy⟩ variable {M : Type*} diff --git a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean index 95813028bae02..d26601f412fd7 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean @@ -85,7 +85,7 @@ instance instLinearOrderedAddCommMonoidWithTopAdditiveOrderDual : LinearOrderedAddCommMonoidWithTop (Additive αᵒᵈ) := { Additive.orderedAddCommMonoid, Additive.linearOrder with top := (0 : α) - top_add' := fun a ↦ zero_mul (Additive.toMul a) + top_add' := fun a ↦ zero_mul a.toMul le_top := fun _ ↦ zero_le' } variable [NoZeroDivisors α] @@ -195,7 +195,7 @@ instance : LinearOrderedAddCommGroupWithTop (Additive αᵒᵈ) := { Additive.subNegMonoid, instLinearOrderedAddCommMonoidWithTopAdditiveOrderDual, Additive.instNontrivial with neg_top := inv_zero (G₀ := α) - add_neg_cancel := fun a ha ↦ mul_inv_cancel₀ (G₀ := α) (id ha : Additive.toMul a ≠ 0) } + add_neg_cancel := fun a ha ↦ mul_inv_cancel₀ (G₀ := α) (id ha : a.toMul ≠ 0) } @[deprecated pow_lt_pow_right₀ (since := "2024-11-18")] lemma pow_lt_pow_succ (ha : 1 < a) : a ^ n < a ^ n.succ := pow_lt_pow_right₀ ha n.lt_succ_self @@ -219,7 +219,7 @@ theorem ofAdd_toDual_eq_zero_iff [LinearOrderedAddCommMonoidWithTop α] @[simp] theorem ofDual_toAdd_eq_top_iff [LinearOrderedAddCommMonoidWithTop α] - (x : Multiplicative αᵒᵈ) : OrderDual.ofDual (Multiplicative.toAdd x) = ⊤ ↔ x = 0 := Iff.rfl + (x : Multiplicative αᵒᵈ) : OrderDual.ofDual x.toAdd = ⊤ ↔ x = 0 := Iff.rfl @[simp] theorem ofAdd_bot [LinearOrderedAddCommMonoidWithTop α] : @@ -227,7 +227,7 @@ theorem ofAdd_bot [LinearOrderedAddCommMonoidWithTop α] : @[simp] theorem ofDual_toAdd_zero [LinearOrderedAddCommMonoidWithTop α] : - OrderDual.ofDual (Multiplicative.toAdd (0 : Multiplicative αᵒᵈ)) = ⊤ := rfl + OrderDual.ofDual (0 : Multiplicative αᵒᵈ).toAdd = ⊤ := rfl instance [LinearOrderedAddCommGroupWithTop α] : LinearOrderedCommGroupWithZero (Multiplicative αᵒᵈ) := diff --git a/Mathlib/Algebra/Order/Monoid/ToMulBot.lean b/Mathlib/Algebra/Order/Monoid/ToMulBot.lean index ab578ee5f849e..79f4fd7e4a107 100644 --- a/Mathlib/Algebra/Order/Monoid/ToMulBot.lean +++ b/Mathlib/Algebra/Order/Monoid/ToMulBot.lean @@ -33,7 +33,7 @@ theorem toMulBot_zero : toMulBot (0 : WithZero (Multiplicative α)) = Multiplica @[simp] theorem toMulBot_coe (x : Multiplicative α) : - toMulBot ↑x = Multiplicative.ofAdd (↑(Multiplicative.toAdd x) : WithBot α) := + toMulBot ↑x = Multiplicative.ofAdd (↑x.toAdd : WithBot α) := rfl @[simp] diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean index a975f6b4f3384..8a7c336d6901c 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean @@ -79,11 +79,11 @@ theorem ofMul_lt {a b : α} : ofMul a < ofMul b ↔ a < b := Iff.rfl @[simp] -theorem toMul_le {a b : Additive α} : toMul a ≤ toMul b ↔ a ≤ b := +theorem toMul_le {a b : Additive α} : a.toMul ≤ b.toMul ↔ a ≤ b := Iff.rfl @[simp] -theorem toMul_lt {a b : Additive α} : toMul a < toMul b ↔ a < b := +theorem toMul_lt {a b : Additive α} : a.toMul < b.toMul ↔ a < b := Iff.rfl end Additive @@ -101,11 +101,11 @@ theorem ofAdd_lt {a b : α} : ofAdd a < ofAdd b ↔ a < b := Iff.rfl @[simp] -theorem toAdd_le {a b : Multiplicative α} : toAdd a ≤ toAdd b ↔ a ≤ b := +theorem toAdd_le {a b : Multiplicative α} : a.toAdd ≤ b.toAdd ↔ a ≤ b := Iff.rfl @[simp] -theorem toAdd_lt {a b : Multiplicative α} : toAdd a < toAdd b ↔ a < b := +theorem toAdd_lt {a b : Multiplicative α} : a.toAdd < b.toAdd ↔ a < b := Iff.rfl end Multiplicative diff --git a/Mathlib/Algebra/Order/SuccPred/TypeTags.lean b/Mathlib/Algebra/Order/SuccPred/TypeTags.lean index 6cda0bda7a3e5..4ee0095a26d7c 100644 --- a/Mathlib/Algebra/Order/SuccPred/TypeTags.lean +++ b/Mathlib/Algebra/Order/SuccPred/TypeTags.lean @@ -37,18 +37,18 @@ open Additive Multiplicative @[simp] lemma succ_ofMul [Preorder X] [SuccOrder X] (x : X) : succ (ofMul x) = ofMul (succ x) := rfl @[simp] lemma succ_toMul [Preorder X] [SuccOrder X] (x : Additive X) : - succ (toMul x) = toMul (succ x) := rfl + succ x.toMul = (succ x).toMul := rfl @[simp] lemma succ_ofAdd [Preorder X] [SuccOrder X] (x : X) : succ (ofAdd x) = ofAdd (succ x) := rfl @[simp] lemma succ_toAdd [Preorder X] [SuccOrder X] (x : Multiplicative X) : - succ (toAdd x) = toAdd (succ x) := rfl + succ x.toAdd = (succ x).toAdd := rfl @[simp] lemma pred_ofMul [Preorder X] [PredOrder X] (x : X) : pred (ofMul x) = ofMul (pred x) := rfl @[simp] lemma pred_toMul [Preorder X] [PredOrder X] (x : Additive X) : - pred (toMul x) = toMul (pred x) := rfl + pred x.toMul = (pred x).toMul := rfl @[simp] lemma pred_ofAdd [Preorder X] [PredOrder X] (x : X) : pred (ofAdd x) = ofAdd (pred x) := rfl @[simp] lemma pred_toAdd [Preorder X] [PredOrder X] (x : Multiplicative X) : - pred (toAdd x) = toAdd (pred x) := rfl + pred x.toAdd = (pred x).toAdd := rfl end Order diff --git a/Mathlib/Analysis/Normed/Group/Constructions.lean b/Mathlib/Analysis/Normed/Group/Constructions.lean index c9ee9cc5d178f..95eb19382e965 100644 --- a/Mathlib/Analysis/Normed/Group/Constructions.lean +++ b/Mathlib/Analysis/Normed/Group/Constructions.lean @@ -103,11 +103,11 @@ variable [Norm E] instance Additive.toNorm : Norm (Additive E) := ‹Norm E› instance Multiplicative.toNorm : Norm (Multiplicative E) := ‹Norm E› -@[simp] lemma norm_toMul (x) : ‖(toMul x : E)‖ = ‖x‖ := rfl +@[simp] lemma norm_toMul (x : Additive E) : ‖(x.toMul : E)‖ = ‖x‖ := rfl @[simp] lemma norm_ofMul (x : E) : ‖ofMul x‖ = ‖x‖ := rfl -@[simp] lemma norm_toAdd (x) : ‖(toAdd x : E)‖ = ‖x‖ := rfl +@[simp] lemma norm_toAdd (x : Multiplicative E) : ‖(x.toAdd : E)‖ = ‖x‖ := rfl @[simp] lemma norm_ofAdd (x : E) : ‖ofAdd x‖ = ‖x‖ := rfl @@ -120,23 +120,23 @@ instance Additive.toNNNorm : NNNorm (Additive E) := ‹NNNorm E› instance Multiplicative.toNNNorm : NNNorm (Multiplicative E) := ‹NNNorm E› -@[simp] lemma nnnorm_toMul (x) : ‖(toMul x : E)‖₊ = ‖x‖₊ := rfl +@[simp] lemma nnnorm_toMul (x : Additive E) : ‖(x.toMul : E)‖₊ = ‖x‖₊ := rfl @[simp] lemma nnnorm_ofMul (x : E) : ‖ofMul x‖₊ = ‖x‖₊ := rfl -@[simp] lemma nnnorm_toAdd (x) : ‖(toAdd x : E)‖₊ = ‖x‖₊ := rfl +@[simp] lemma nnnorm_toAdd (x : Multiplicative E) : ‖(x.toAdd : E)‖₊ = ‖x‖₊ := rfl @[simp] lemma nnnorm_ofAdd (x : E) : ‖ofAdd x‖₊ = ‖x‖₊ := rfl end NNNorm instance Additive.seminormedAddGroup [SeminormedGroup E] : SeminormedAddGroup (Additive E) where - dist_eq x y := dist_eq_norm_div (toMul x) (toMul y) + dist_eq x y := dist_eq_norm_div x.toMul y.toMul instance Multiplicative.seminormedGroup [SeminormedAddGroup E] : SeminormedGroup (Multiplicative E) where - dist_eq x y := dist_eq_norm_sub (toMul x) (toMul y) + dist_eq x y := dist_eq_norm_sub x.toAdd y.toAdd instance Additive.seminormedCommGroup [SeminormedCommGroup E] : SeminormedAddCommGroup (Additive E) := diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index 1054c7a046c4b..f789146aa2884 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -188,7 +188,7 @@ theorem exp_add : exp (x + y) = exp x * exp y := by /-- the exponential function as a monoid hom from `Multiplicative ℂ` to `ℂ` -/ @[simps] noncomputable def expMonoidHom : MonoidHom (Multiplicative ℂ) ℂ := - { toFun := fun z => exp (Multiplicative.toAdd z), + { toFun := fun z => exp z.toAdd, map_one' := by simp, map_mul' := by simp [exp_add] } @@ -689,7 +689,7 @@ nonrec theorem exp_add : exp (x + y) = exp x * exp y := by simp [exp_add, exp] /-- the exponential function as a monoid hom from `Multiplicative ℝ` to `ℝ` -/ @[simps] noncomputable def expMonoidHom : MonoidHom (Multiplicative ℝ) ℝ := - { toFun := fun x => exp (Multiplicative.toAdd x), + { toFun := fun x => exp x.toAdd, map_one' := by simp, map_mul' := by simp [exp_add] } diff --git a/Mathlib/Data/Int/Cast/Lemmas.lean b/Mathlib/Data/Int/Cast/Lemmas.lean index edf29c1a81e49..d887d52523d54 100644 --- a/Mathlib/Data/Int/Cast/Lemmas.lean +++ b/Mathlib/Data/Int/Cast/Lemmas.lean @@ -295,14 +295,14 @@ lemma zmultiplesHom_apply (x : β) (n : ℤ) : zmultiplesHom β x n = n • x := lemma zmultiplesHom_symm_apply (f : ℤ →+ β) : (zmultiplesHom β).symm f = f 1 := rfl @[to_additive existing (attr := simp)] -lemma zpowersHom_apply (x : α) (n : Multiplicative ℤ) : zpowersHom α x n = x ^ toAdd n := rfl +lemma zpowersHom_apply (x : α) (n : Multiplicative ℤ) : zpowersHom α x n = x ^ n.toAdd := rfl @[to_additive existing (attr := simp)] lemma zpowersHom_symm_apply (f : Multiplicative ℤ →* α) : (zpowersHom α).symm f = f (ofAdd 1) := rfl lemma MonoidHom.apply_mint (f : Multiplicative ℤ →* α) (n : Multiplicative ℤ) : - f n = f (ofAdd 1) ^ (toAdd n) := by + f n = f (ofAdd 1) ^ n.toAdd := by rw [← zpowersHom_symm_apply, ← zpowersHom_apply, Equiv.apply_symm_apply] lemma AddMonoidHom.apply_int (f : ℤ →+ β) (n : ℤ) : f n = n • f 1 := by @@ -324,7 +324,7 @@ def zpowersMulHom : α ≃* (Multiplicative ℤ →* α) := variable {α} @[simp] -lemma zpowersMulHom_apply (x : α) (n : Multiplicative ℤ) : zpowersMulHom α x n = x ^ toAdd n := rfl +lemma zpowersMulHom_apply (x : α) (n : Multiplicative ℤ) : zpowersMulHom α x n = x ^ n.toAdd := rfl @[simp] lemma zpowersMulHom_symm_apply (f : Multiplicative ℤ →* α) : diff --git a/Mathlib/Data/Int/WithZero.lean b/Mathlib/Data/Int/WithZero.lean index 5cf5aebd6da66..de4c3c4143538 100644 --- a/Mathlib/Data/Int/WithZero.lean +++ b/Mathlib/Data/Int/WithZero.lean @@ -14,7 +14,7 @@ the morphism `WithZeroMultInt.toNNReal`. ## Main Definitions * `WithZeroMultInt.toNNReal` : The `MonoidWithZeroHom` from `ℤₘ₀ → ℝ≥0` sending `0 ↦ 0` and - `x ↦ e^(Multiplicative.toAdd (WithZero.unzero hx)` when `x ≠ 0`, for a nonzero `e : ℝ≥0`. + `x ↦ e^((WithZero.unzero hx).toAdd)` when `x ≠ 0`, for a nonzero `e : ℝ≥0`. ## Main Results @@ -37,9 +37,9 @@ open Multiplicative WithZero namespace WithZeroMulInt /-- Given a nonzero `e : ℝ≥0`, this is the map `ℤₘ₀ → ℝ≥0` sending `0 ↦ 0` and - `x ↦ e^(Multiplicative.toAdd (WithZero.unzero hx)` when `x ≠ 0` as a `MonoidWithZeroHom`. -/ + `x ↦ e^(WithZero.unzero hx).toAdd` when `x ≠ 0` as a `MonoidWithZeroHom`. -/ def toNNReal {e : ℝ≥0} (he : e ≠ 0) : ℤₘ₀ →*₀ ℝ≥0 where - toFun := fun x ↦ if hx : x = 0 then 0 else e ^ Multiplicative.toAdd (WithZero.unzero hx) + toFun := fun x ↦ if hx : x = 0 then 0 else e ^ (WithZero.unzero hx).toAdd map_zero' := rfl map_one' := by simp only [dif_neg one_ne_zero] @@ -63,7 +63,7 @@ theorem toNNReal_pos_apply {e : ℝ≥0} (he : e ≠ 0) {x : ℤₘ₀} (hx : x split_ifs; rfl theorem toNNReal_neg_apply {e : ℝ≥0} (he : e ≠ 0) {x : ℤₘ₀} (hx : x ≠ 0) : - toNNReal he x = e ^ Multiplicative.toAdd (WithZero.unzero hx) := by + toNNReal he x = e ^ (WithZero.unzero hx).toAdd := by simp only [toNNReal, MonoidWithZeroHom.coe_mk, ZeroHom.coe_mk] split_ifs · tauto diff --git a/Mathlib/Data/Nat/Cast/Basic.lean b/Mathlib/Data/Nat/Cast/Basic.lean index d9fd72598d7bf..63025e62d2769 100644 --- a/Mathlib/Data/Nat/Cast/Basic.lean +++ b/Mathlib/Data/Nat/Cast/Basic.lean @@ -222,7 +222,7 @@ lemma multiplesHom_apply (x : β) (n : ℕ) : multiplesHom β x n = n • x := r @[to_additive existing (attr := simp)] lemma powersHom_apply (x : α) (n : Multiplicative ℕ) : - powersHom α x n = x ^ Multiplicative.toAdd n := rfl + powersHom α x n = x ^ n.toAdd := rfl lemma multiplesHom_symm_apply (f : ℕ →+ β) : (multiplesHom β).symm f = f 1 := rfl @@ -231,7 +231,7 @@ lemma powersHom_symm_apply (f : Multiplicative ℕ →* α) : (powersHom α).symm f = f (Multiplicative.ofAdd 1) := rfl lemma MonoidHom.apply_mnat (f : Multiplicative ℕ →* α) (n : Multiplicative ℕ) : - f n = f (Multiplicative.ofAdd 1) ^ (Multiplicative.toAdd n) := by + f n = f (Multiplicative.ofAdd 1) ^ n.toAdd := by rw [← powersHom_symm_apply, ← powersHom_apply, Equiv.apply_symm_apply] @[ext] @@ -258,7 +258,7 @@ def powersMulHom : α ≃* (Multiplicative ℕ →* α) := @[simp] lemma multiplesAddHom_apply (x : β) (n : ℕ) : multiplesAddHom β x n = n • x := rfl @[simp] -lemma powersMulHom_apply (x : α) (n : Multiplicative ℕ) : powersMulHom α x n = x ^ toAdd n := rfl +lemma powersMulHom_apply (x : α) (n : Multiplicative ℕ) : powersMulHom α x n = x ^ n.toAdd := rfl @[simp] lemma multiplesAddHom_symm_apply (f : ℕ →+ β) : (multiplesAddHom β).symm f = f 1 := rfl diff --git a/Mathlib/Data/ZMod/IntUnitsPower.lean b/Mathlib/Data/ZMod/IntUnitsPower.lean index 3927a7e59bb96..4b5cace9f38cd 100644 --- a/Mathlib/Data/ZMod/IntUnitsPower.lean +++ b/Mathlib/Data/ZMod/IntUnitsPower.lean @@ -24,7 +24,7 @@ by using `Module R (Additive M)` in its place, especially since this already has -/ instance : SMul (ZMod 2) (Additive ℤˣ) where - smul z au := .ofMul <| Additive.toMul au ^ z.val + smul z au := .ofMul <| au.toMul ^ z.val lemma ZMod.smul_units_def (z : ZMod 2) (au : Additive ℤˣ) : z • au = z.val • au := rfl @@ -34,7 +34,7 @@ lemma ZMod.natCast_smul_units (n : ℕ) (au : Additive ℤˣ) : (n : ZMod 2) • /-- This is an indirect way of saying that `ℤˣ` has a power operation by `ZMod 2`. -/ instance : Module (ZMod 2) (Additive ℤˣ) where - smul z au := .ofMul <| Additive.toMul au ^ z.val + smul z au := .ofMul <| au.toMul ^ z.val one_smul _ := Additive.toMul.injective <| pow_one _ mul_smul z₁ z₂ au := Additive.toMul.injective <| by dsimp only [ZMod.smul_units_def, toMul_nsmul] @@ -44,7 +44,7 @@ instance : Module (ZMod 2) (Additive ℤˣ) where add_smul z₁ z₂ au := Additive.toMul.injective <| by dsimp only [ZMod.smul_units_def, toMul_nsmul, toMul_add] rw [← pow_add, ZMod.val_add, ← Int.units_pow_eq_pow_mod_two] - zero_smul au := Additive.toMul.injective <| pow_zero (Additive.toMul au) + zero_smul au := Additive.toMul.injective <| pow_zero au.toMul section CommSemiring variable {R : Type*} [CommSemiring R] [Module R (Additive ℤˣ)] @@ -55,7 +55,7 @@ In lemma names, this operations is called `uzpow` to match `zpow`. Notably this is satisfied by `R ∈ {ℕ, ℤ, ZMod 2}`. -/ instance Int.instUnitsPow : Pow ℤˣ R where - pow u r := Additive.toMul (r • Additive.ofMul u) + pow u r := (r • Additive.ofMul u).toMul -- The above instances form no typeclass diamonds with the standard power operators -- but we will need `reducible_and_instances` which currently fails https://github.com/leanprover-community/mathlib4/issues/10906 @@ -65,10 +65,10 @@ example : Int.instUnitsPow = DivInvMonoid.Pow := rfl @[simp] lemma ofMul_uzpow (u : ℤˣ) (r : R) : Additive.ofMul (u ^ r) = r • Additive.ofMul u := rfl @[simp] lemma toMul_uzpow (u : Additive ℤˣ) (r : R) : - Additive.toMul (r • u) = Additive.toMul u ^ r := rfl + (r • u).toMul = u.toMul ^ r := rfl @[norm_cast] lemma uzpow_natCast (u : ℤˣ) (n : ℕ) : u ^ (n : R) = u ^ n := by - change Additive.toMul ((n : R) • Additive.ofMul u) = _ + change ((n : R) • Additive.ofMul u).toMul = _ rw [Nat.cast_smul_eq_nsmul, toMul_nsmul, toMul_ofMul] -- See note [no_index around OfNat.ofNat] @@ -106,7 +106,7 @@ lemma uzpow_neg (s : ℤˣ) (x : R) : s ^ (-x) = (s ^ x)⁻¹ := Additive.ofMul.injective <| neg_smul x (Additive.ofMul s) @[norm_cast] lemma uzpow_intCast (u : ℤˣ) (z : ℤ) : u ^ (z : R) = u ^ z := by - change Additive.toMul ((z : R) • Additive.ofMul u) = _ + change ((z : R) • Additive.ofMul u).toMul = _ rw [Int.cast_smul_eq_zsmul, toMul_zsmul, toMul_ofMul] end CommRing diff --git a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean index 0b03b2f6c460d..e403b3b59936a 100644 --- a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean +++ b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean @@ -253,7 +253,7 @@ theorem commute_iff_commute {f g : CircleDeg1Lift} : Commute f g ↔ Function.Co `translation (Multiplicative.ofAdd x)`. -/ def translate : Multiplicative ℝ →* CircleDeg1Liftˣ := MonoidHom.toHomUnits <| { toFun := fun x => - ⟨⟨fun y => Multiplicative.toAdd x + y, fun _ _ h => add_le_add_left h _⟩, fun _ => + ⟨⟨fun y => x.toAdd + y, fun _ _ h => add_le_add_left h _⟩, fun _ => (add_assoc _ _ _).symm⟩ map_one' := ext <| zero_add map_mul' := fun _ _ => ext <| add_assoc _ _ } diff --git a/Mathlib/GroupTheory/FiniteAbelian/Duality.lean b/Mathlib/GroupTheory/FiniteAbelian/Duality.lean index 308c47949bf39..dafdda97bea03 100644 --- a/Mathlib/GroupTheory/FiniteAbelian/Duality.lean +++ b/Mathlib/GroupTheory/FiniteAbelian/Duality.lean @@ -42,9 +42,9 @@ lemma exists_apply_ne_one_aux (H : ∀ n : ℕ, n ∣ Monoid.exponent G → ∀ obtain ⟨i, hi⟩ : ∃ i : ι, e a i ≠ 1 := by contrapose! ha exact (MulEquiv.map_eq_one_iff e).mp <| funext ha - have hi : Multiplicative.toAdd (e a i) ≠ 0 := by + have hi : (e a i).toAdd ≠ 0 := by simp only [ne_eq, toAdd_eq_zero, hi, not_false_eq_true] - obtain ⟨φi, hφi⟩ := H (n i) (dvd_exponent e i) (Multiplicative.toAdd <| e a i) hi + obtain ⟨φi, hφi⟩ := H (n i) (dvd_exponent e i) ((e a i).toAdd) hi use (φi.comp (Pi.evalMonoidHom (fun (i : ι) ↦ Multiplicative (ZMod (n i))) i)).comp e simpa only [coe_comp, coe_coe, Function.comp_apply, Pi.evalMonoidHom_apply, ne_eq] using hφi diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean index 20adcbf583239..63e23ebbb7c77 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean @@ -422,7 +422,7 @@ theorem constVAdd_symm (v : V₁) : (constVAdd k P₁ v).symm = constVAdd k P₁ /-- A more bundled version of `AffineEquiv.constVAdd`. -/ @[simps] def constVAddHom : Multiplicative V₁ →* P₁ ≃ᵃ[k] P₁ where - toFun v := constVAdd k P₁ (Multiplicative.toAdd v) + toFun v := constVAdd k P₁ v.toAdd map_one' := constVAdd_zero _ _ map_mul' := constVAdd_add _ P₁ diff --git a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean index 31c16252e0c0d..4fbfe52301fe9 100644 --- a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean +++ b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean @@ -75,7 +75,7 @@ variable (K) /-- The logarithmic embedding of the units (seen as an `Additive` group). -/ def _root_.NumberField.Units.logEmbedding : Additive ((𝓞 K)ˣ) →+ ({w : InfinitePlace K // w ≠ w₀} → ℝ) := -{ toFun := fun x w => mult w.val * Real.log (w.val ↑(Additive.toMul x)) +{ toFun := fun x w => mult w.val * Real.log (w.val ↑x.toMul) map_zero' := by simp; rfl map_add' := fun _ _ => by simp [Real.log_mul, mul_add]; rfl } @@ -465,7 +465,7 @@ def basisUnitLattice : Basis (Fin (rank K)) ℤ (unitLattice K) := units in `basisModTorsion`. -/ def fundSystem : Fin (rank K) → (𝓞 K)ˣ := -- `:)` prevents the `⧸` decaying to a quotient by `leftRel` when we unfold this later - fun i => Quotient.out (Additive.toMul (basisModTorsion K i):) + fun i => Quotient.out ((basisModTorsion K i).toMul:) theorem fundSystem_mk (i : Fin (rank K)) : Additive.ofMul (QuotientGroup.mk (fundSystem K i)) = (basisModTorsion K i) := by diff --git a/Mathlib/RepresentationTheory/Basic.lean b/Mathlib/RepresentationTheory/Basic.lean index fbac878c2ff51..f17cfb213b7e5 100644 --- a/Mathlib/RepresentationTheory/Basic.lean +++ b/Mathlib/RepresentationTheory/Basic.lean @@ -292,7 +292,7 @@ def ofMulDistribMulAction : Representation ℤ M (Additive G) := ((monoidEndToAdditive G : _ →* _).comp (MulDistribMulAction.toMonoidEnd M G)) @[simp] theorem ofMulDistribMulAction_apply_apply (g : M) (a : Additive G) : - ofMulDistribMulAction M G g a = Additive.ofMul (g • Additive.toMul a) := rfl + ofMulDistribMulAction M G g a = Additive.ofMul (g • a.toMul) := rfl end MulDistribMulAction section Group diff --git a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean index b05641a81a752..24d8d8f5bd37c 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean @@ -710,7 +710,7 @@ theorem H1LequivOfIsTrivial_comp_H1_π [A.IsTrivial] : @[simp] theorem H1LequivOfIsTrivial_H1_π_apply_apply [A.IsTrivial] (f : oneCocycles A) (x : Additive G) : - H1LequivOfIsTrivial A (H1_π A f) x = f (Additive.toMul x) := rfl + H1LequivOfIsTrivial A (H1_π A f) x = f x.toMul := rfl @[simp] theorem H1LequivOfIsTrivial_symm_apply [A.IsTrivial] (f : Additive G →+ A) : (H1LequivOfIsTrivial A).symm f = H1_π A ((oneCocyclesLequivOfIsTrivial A).symm f) := diff --git a/Mathlib/RepresentationTheory/Rep.lean b/Mathlib/RepresentationTheory/Rep.lean index d4ddfa12b4c67..0a14328b88a96 100644 --- a/Mathlib/RepresentationTheory/Rep.lean +++ b/Mathlib/RepresentationTheory/Rep.lean @@ -254,7 +254,7 @@ variable (M G : Type) [Monoid M] [CommGroup G] [MulDistribMulAction M G] def ofMulDistribMulAction : Rep ℤ M := Rep.of (Representation.ofMulDistribMulAction M G) @[simp] theorem ofMulDistribMulAction_ρ_apply_apply (g : M) (a : Additive G) : - (ofMulDistribMulAction M G).ρ g a = Additive.ofMul (g • Additive.toMul a) := rfl + (ofMulDistribMulAction M G).ρ g a = Additive.ofMul (g • a.toMul) := rfl /-- Given an `R`-algebra `S`, the `ℤ`-linear representation associated to the natural action of `S ≃ₐ[R] S` on `Sˣ`. -/ diff --git a/Mathlib/RingTheory/FreeCommRing.lean b/Mathlib/RingTheory/FreeCommRing.lean index 6efc6c11c1cd6..e1b0736aa1404 100644 --- a/Mathlib/RingTheory/FreeCommRing.lean +++ b/Mathlib/RingTheory/FreeCommRing.lean @@ -134,7 +134,7 @@ private def liftToMultiset : (α → R) ≃ (Multiplicative (Multiset α) →* R left_inv f := funext fun x => show (Multiset.map f {x}).prod = _ by simp right_inv F := MonoidHom.ext fun x => let F' := MonoidHom.toAdditive'' F - let x' := Multiplicative.toAdd x + let x' := x.toAdd show (Multiset.map (fun a => F' {a}) x').sum = F' x' by erw [← Multiset.map_map (fun x => F' x) (fun x => {x}), ← AddMonoidHom.map_multiset_sum] exact DFunLike.congr_arg F (Multiset.sum_map_singleton x') diff --git a/Mathlib/RingTheory/LaurentSeries.lean b/Mathlib/RingTheory/LaurentSeries.lean index f8ee85f00a0c0..1e09014fa6640 100644 --- a/Mathlib/RingTheory/LaurentSeries.lean +++ b/Mathlib/RingTheory/LaurentSeries.lean @@ -869,7 +869,7 @@ theorem Cauchy.eventually_mem_nhds {ℱ : Filter K⸨X⸩} (hℱ : Cauchy ℱ) obtain ⟨γ, hU₁⟩ := Valued.mem_nhds.mp hU suffices ∀ᶠ f in ℱ, f ∈ {y : K⸨X⸩ | Valued.v (y - limit hℱ) < ↑γ} by apply this.mono fun _ hf ↦ hU₁ hf - set D := -(Multiplicative.toAdd (WithZero.unzero γ.ne_zero) - 1) with hD₀ + set D := -((WithZero.unzero γ.ne_zero).toAdd - 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] diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index 7d92b1e682ec5..280967dfdc29d 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -115,9 +115,9 @@ theorem nhds_ofMul (x : X) : 𝓝 (ofMul x) = map ofMul (𝓝 x) := rfl theorem nhds_ofAdd (x : X) : 𝓝 (ofAdd x) = map ofAdd (𝓝 x) := rfl -theorem nhds_toMul (x : Additive X) : 𝓝 (toMul x) = map toMul (𝓝 x) := rfl +theorem nhds_toMul (x : Additive X) : 𝓝 x.toMul = map toMul (𝓝 x) := rfl -theorem nhds_toAdd (x : Multiplicative X) : 𝓝 (toAdd x) = map toAdd (𝓝 x) := rfl +theorem nhds_toAdd (x : Multiplicative X) : 𝓝 x.toAdd = map toAdd (𝓝 x) := rfl end diff --git a/Mathlib/Topology/EMetricSpace/Defs.lean b/Mathlib/Topology/EMetricSpace/Defs.lean index 445f6803e42d1..b791062fbf650 100644 --- a/Mathlib/Topology/EMetricSpace/Defs.lean +++ b/Mathlib/Topology/EMetricSpace/Defs.lean @@ -657,11 +657,11 @@ theorem edist_ofAdd (a b : X) : edist (ofAdd a) (ofAdd b) = edist a b := rfl @[simp] -theorem edist_toMul (a b : Additive X) : edist (toMul a) (toMul b) = edist a b := +theorem edist_toMul (a b : Additive X) : edist a.toMul b.toMul = edist a b := rfl @[simp] -theorem edist_toAdd (a b : Multiplicative X) : edist (toAdd a) (toAdd b) = edist a b := +theorem edist_toAdd (a b : Multiplicative X) : edist a.toAdd b.toAdd = edist a b := rfl end diff --git a/Mathlib/Topology/MetricSpace/Defs.lean b/Mathlib/Topology/MetricSpace/Defs.lean index 93866859fadb0..e678bdcfedb6b 100644 --- a/Mathlib/Topology/MetricSpace/Defs.lean +++ b/Mathlib/Topology/MetricSpace/Defs.lean @@ -197,9 +197,9 @@ instance : Dist (Multiplicative X) := ‹Dist X› @[simp] theorem dist_ofAdd (a b : X) : dist (ofAdd a) (ofAdd b) = dist a b := rfl -@[simp] theorem dist_toMul (a b : Additive X) : dist (toMul a) (toMul b) = dist a b := rfl +@[simp] theorem dist_toMul (a b : Additive X) : dist a.toMul b.toMul = dist a b := rfl -@[simp] theorem dist_toAdd (a b : Multiplicative X) : dist (toAdd a) (toAdd b) = dist a b := rfl +@[simp] theorem dist_toAdd (a b : Multiplicative X) : dist a.toAdd b.toAdd = dist a b := rfl end @@ -211,10 +211,10 @@ variable [PseudoMetricSpace X] @[simp] theorem nndist_ofAdd (a b : X) : nndist (ofAdd a) (ofAdd b) = nndist a b := rfl -@[simp] theorem nndist_toMul (a b : Additive X) : nndist (toMul a) (toMul b) = nndist a b := rfl +@[simp] theorem nndist_toMul (a b : Additive X) : nndist a.toMul b.toMul = nndist a b := rfl @[simp] -theorem nndist_toAdd (a b : Multiplicative X) : nndist (toAdd a) (toAdd b) = nndist a b := rfl +theorem nndist_toAdd (a b : Multiplicative X) : nndist a.toAdd b.toAdd = nndist a b := rfl end diff --git a/Mathlib/Topology/MetricSpace/IsometricSMul.lean b/Mathlib/Topology/MetricSpace/IsometricSMul.lean index 8bbc25f34f0ab..7af5ce1593cca 100644 --- a/Mathlib/Topology/MetricSpace/IsometricSMul.lean +++ b/Mathlib/Topology/MetricSpace/IsometricSMul.lean @@ -420,26 +420,26 @@ instance Pi.isometricSMul'' {ι} {M : ι → Type*} [Fintype ι] [∀ i, Mul (M ⟨fun c => .piMap (fun i (x : M i) => x * c.unop i) fun _ => isometry_mul_right _⟩ instance Additive.isometricVAdd : IsometricVAdd (Additive M) X := - ⟨fun c => isometry_smul X (toMul c)⟩ + ⟨fun c => isometry_smul X c.toMul⟩ instance Additive.isometricVAdd' [Mul M] [PseudoEMetricSpace M] [IsometricSMul M M] : IsometricVAdd (Additive M) (Additive M) := - ⟨fun c x y => edist_smul_left (toMul c) (toMul x) (toMul y)⟩ + ⟨fun c x y => edist_smul_left c.toMul x.toMul y.toMul⟩ instance Additive.isometricVAdd'' [Mul M] [PseudoEMetricSpace M] [IsometricSMul Mᵐᵒᵖ M] : IsometricVAdd (Additive M)ᵃᵒᵖ (Additive M) := - ⟨fun c x y => edist_smul_left (MulOpposite.op (toMul c.unop)) (toMul x) (toMul y)⟩ + ⟨fun c x y => edist_smul_left (MulOpposite.op c.unop.toMul) x.toMul y.toMul⟩ instance Multiplicative.isometricSMul {M X} [VAdd M X] [PseudoEMetricSpace X] [IsometricVAdd M X] : IsometricSMul (Multiplicative M) X := - ⟨fun c => isometry_vadd X (toAdd c)⟩ + ⟨fun c => isometry_vadd X c.toAdd⟩ instance Multiplicative.isometricSMul' [Add M] [PseudoEMetricSpace M] [IsometricVAdd M M] : IsometricSMul (Multiplicative M) (Multiplicative M) := - ⟨fun c x y => edist_vadd_left (toAdd c) (toAdd x) (toAdd y)⟩ + ⟨fun c x y => edist_vadd_left c.toAdd x.toAdd y.toAdd⟩ instance Multiplicative.isometricVAdd'' [Add M] [PseudoEMetricSpace M] [IsometricVAdd Mᵃᵒᵖ M] : IsometricSMul (Multiplicative M)ᵐᵒᵖ (Multiplicative M) := - ⟨fun c x y => edist_vadd_left (AddOpposite.op (toAdd c.unop)) (toAdd x) (toAdd y)⟩ + ⟨fun c x y => edist_vadd_left (AddOpposite.op c.unop.toAdd) x.toAdd y.toAdd⟩ end Instances From c9cb88e7c1a95c987df4d3d68dd55f4c90d5bf11 Mon Sep 17 00:00:00 2001 From: grunweg Date: Sat, 23 Nov 2024 00:14:53 +0000 Subject: [PATCH 212/829] chore(workflows): add another missing job name (#19386) Analogous to #19382. --- .github/workflows/sync_closed_tasks.yaml | 1 + 1 file changed, 1 insertion(+) diff --git a/.github/workflows/sync_closed_tasks.yaml b/.github/workflows/sync_closed_tasks.yaml index 9c73e7e8a24c6..6c5383d62099c 100644 --- a/.github/workflows/sync_closed_tasks.yaml +++ b/.github/workflows/sync_closed_tasks.yaml @@ -12,6 +12,7 @@ on: jobs: build: + name: "Cross off linked issues" runs-on: ubuntu-latest steps: - name: Cross off any linked issue and PR references From 5483f1cf4b7d47e8843e2f4b71605bc74c60a3a7 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Sat, 23 Nov 2024 02:13:37 +0000 Subject: [PATCH 213/829] chore: update Mathlib dependencies 2024-11-23 (#19389) 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 9d7aff8b1542f..762bb1280e3ba 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "33d7f346440869364a2cd077bde8cebf73243aaa", + "rev": "8b52587ff32e2e443cce6109b5305341289339e7", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From 5b58f54371168aba52c6d88e02af49b56ee41d6f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ga=C3=ABtan=20Serr=C3=A9?= Date: Sat, 23 Nov 2024 04:20:30 +0000 Subject: [PATCH 214/829] feat(Order/Filter/Basic): eventually_iff_all_subsets and specializations (#13275) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The set `{x | p x}` belongs to `f` iff `∀ s, {x | x ∈ s → p x} ∈ f`. Specialization for `EventuallyEq` and `Eventually.LE`. Suggested by @urkud. Co-authored-by: Gaëtan Serré <56162277+gaetanserre@users.noreply.github.com> Co-authored-by: Johan Commelin --- Mathlib/Order/Filter/Basic.lean | 13 +++++++++++++ 1 file changed, 13 insertions(+) diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index ee7fd636ed470..995ddfea3234b 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -895,6 +895,11 @@ theorem eventually_inf_principal {f : Filter α} {p : α → Prop} {s : Set α} (∀ᶠ x in f ⊓ 𝓟 s, p x) ↔ ∀ᶠ x in f, x ∈ s → p x := mem_inf_principal +theorem eventually_iff_all_subsets {f : Filter α} {p : α → Prop} : + (∀ᶠ x in f, p x) ↔ ∀ (s : Set α), ∀ᶠ x in f, x ∈ s → p x where + mp h _ := by filter_upwards [h] with _ pa _ using pa + mpr h := by filter_upwards [h univ] with _ pa using pa (by simp) + /-! ### Frequently -/ theorem Eventually.frequently {f : Filter α} [NeBot f] {p : α → Prop} (h : ∀ᶠ x in f, p x) : @@ -1224,6 +1229,10 @@ theorem eventuallyEq_iff_sub [AddGroup β] {f g : α → β} {l : Filter α} : f =ᶠ[l] g ↔ f - g =ᶠ[l] 0 := ⟨fun h => h.sub_eq, fun h => by simpa using h.add (EventuallyEq.refl l g)⟩ +theorem eventuallyEq_iff_all_subsets {f g : α → β} {l : Filter α} : + f =ᶠ[l] g ↔ ∀ s : Set α, ∀ᶠ x in l, x ∈ s → f x = g x := + eventually_iff_all_subsets + section LE variable [LE β] {l : Filter α} @@ -1236,6 +1245,10 @@ theorem eventuallyLE_congr {f f' g g' : α → β} (hf : f =ᶠ[l] f') (hg : g = f ≤ᶠ[l] g ↔ f' ≤ᶠ[l] g' := ⟨fun H => H.congr hf hg, fun H => H.congr hf.symm hg.symm⟩ +theorem eventuallyLE_iff_all_subsets {f g : α → β} {l : Filter α} : + f ≤ᶠ[l] g ↔ ∀ s : Set α, ∀ᶠ x in l, x ∈ s → f x ≤ g x := + eventually_iff_all_subsets + end LE section Preorder From e360f005e729a625b4f36183a23a57e1865a527d Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Sat, 23 Nov 2024 04:58:45 +0000 Subject: [PATCH 215/829] feat: add `Subgroup.index_antitone` and `Subgroup.index_strictAnti` (#19188) from flt-regular. --- Mathlib/GroupTheory/Index.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/Mathlib/GroupTheory/Index.lean b/Mathlib/GroupTheory/Index.lean index 8357feb56afc9..c31a41f62822f 100644 --- a/Mathlib/GroupTheory/Index.lean +++ b/Mathlib/GroupTheory/Index.lean @@ -554,6 +554,17 @@ variable {H K} theorem finiteIndex_of_le [FiniteIndex H] (h : H ≤ K) : FiniteIndex K := ⟨ne_zero_of_dvd_ne_zero FiniteIndex.finiteIndex (index_dvd_of_le h)⟩ +@[to_additive (attr := gcongr)] +lemma index_antitone (h : H ≤ K) [H.FiniteIndex] : K.index ≤ H.index := + Nat.le_of_dvd (Nat.zero_lt_of_ne_zero FiniteIndex.finiteIndex) (index_dvd_of_le h) + +@[to_additive (attr := gcongr)] +lemma index_strictAnti (h : H < K) [H.FiniteIndex] : K.index < H.index := by + have h0 : K.index ≠ 0 := (finiteIndex_of_le h.le).finiteIndex + apply lt_of_le_of_ne (index_antitone h.le) + rw [← relindex_mul_index h.le, Ne, eq_comm, mul_eq_right₀ h0, relindex_eq_one] + exact h.not_le + variable (H K) @[to_additive] From 506763e8dc880a44e92ebe195c4d369d3dcef5f0 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Sat, 23 Nov 2024 04:58:46 +0000 Subject: [PATCH 216/829] chore: rename *tower_top_of_injective lemmas (#19246) --- Mathlib/FieldTheory/IsAlgClosed/Basic.lean | 2 +- Mathlib/RingTheory/Algebraic.lean | 38 +++++++++++-------- Mathlib/RingTheory/AlgebraicIndependent.lean | 3 +- Mathlib/RingTheory/Localization/Integral.lean | 2 +- 4 files changed, 25 insertions(+), 20 deletions(-) diff --git a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean index 2ec4275183cac..c2bc06cd1f067 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean @@ -410,7 +410,7 @@ noncomputable def equivOfEquivAux (hSR : S ≃+* R) : haveI : IsScalarTower S R L := IsScalarTower.of_algebraMap_eq (by simp [RingHom.algebraMap_toAlgebra]) haveI : NoZeroSMulDivisors R S := NoZeroSMulDivisors.of_algebraMap_injective hSR.symm.injective - have : Algebra.IsAlgebraic R L := (IsAlgClosure.isAlgebraic.tower_top_of_injective + have : Algebra.IsAlgebraic R L := (IsAlgClosure.isAlgebraic.extendScalars (show Function.Injective (algebraMap S R) from hSR.injective)) refine ⟨equivOfAlgebraic' R S L M, ?_⟩ diff --git a/Mathlib/RingTheory/Algebraic.lean b/Mathlib/RingTheory/Algebraic.lean index e8fe139139573..4752101e3862d 100644 --- a/Mathlib/RingTheory/Algebraic.lean +++ b/Mathlib/RingTheory/Algebraic.lean @@ -499,31 +499,34 @@ variable [Algebra R S] [Algebra S A] [Algebra R A] [IsScalarTower R S A] /-- If `x` is algebraic over `R`, then `x` is algebraic over `S` when `S` is an extension of `R`, and the map from `R` to `S` is injective. -/ -theorem IsAlgebraic.tower_top_of_injective - (hinj : Function.Injective (algebraMap R S)) {x : A} +theorem IsAlgebraic.extendScalars (hinj : Function.Injective (algebraMap R S)) {x : A} (A_alg : IsAlgebraic R x) : IsAlgebraic S x := let ⟨p, hp₁, hp₂⟩ := A_alg ⟨p.map (algebraMap _ _), by rwa [Ne, ← degree_eq_bot, degree_map_eq_of_injective hinj, degree_eq_bot], by simpa⟩ -/-- A special case of `IsAlgebraic.tower_top_of_injective`. This is extracted as a theorem - because in some cases `IsAlgebraic.tower_top_of_injective` will just runs out of memory. -/ +@[deprecated (since := "2024-11-18")] +alias IsAlgebraic.tower_top_of_injective := IsAlgebraic.extendScalars + +/-- A special case of `IsAlgebraic.extendScalars`. This is extracted as a theorem + because in some cases `IsAlgebraic.extendScalars` will just runs out of memory. -/ theorem IsAlgebraic.tower_top_of_subalgebra_le {A B : Subalgebra R S} (hle : A ≤ B) {x : S} (h : IsAlgebraic A x) : IsAlgebraic B x := by letI : Algebra A B := (Subalgebra.inclusion hle).toAlgebra haveI : IsScalarTower A B S := .of_algebraMap_eq fun _ ↦ rfl - exact h.tower_top_of_injective (Subalgebra.inclusion_injective hle) + exact h.extendScalars (Subalgebra.inclusion_injective hle) /-- If `x` is transcendental over `S`, then `x` is transcendental over `R` when `S` is an extension of `R`, and the map from `R` to `S` is injective. -/ -theorem Transcendental.of_tower_top_of_injective - (hinj : Function.Injective (algebraMap R S)) {x : A} - (h : Transcendental S x) : Transcendental R x := - fun H ↦ h (H.tower_top_of_injective hinj) +theorem Transcendental.restrictScalars (hinj : Function.Injective (algebraMap R S)) {x : A} + (h : Transcendental S x) : Transcendental R x := fun H ↦ h (H.extendScalars hinj) + +@[deprecated (since := "2024-11-18")] +alias Transcendental.of_tower_top_of_injective := Transcendental.restrictScalars -/-- A special case of `Transcendental.of_tower_top_of_injective`. This is extracted as a theorem - because in some cases `Transcendental.of_tower_top_of_injective` will just runs out of memory. -/ +/-- A special case of `Transcendental.restrictScalars`. This is extracted as a theorem + because in some cases `Transcendental.restrictScalars` will just runs out of memory. -/ theorem Transcendental.of_tower_top_of_subalgebra_le {A B : Subalgebra R S} (hle : A ≤ B) {x : S} (h : Transcendental B x) : Transcendental A x := @@ -531,9 +534,12 @@ theorem Transcendental.of_tower_top_of_subalgebra_le /-- If A is an algebraic algebra over R, then A is algebraic over S when S is an extension of R, and the map from `R` to `S` is injective. -/ -theorem Algebra.IsAlgebraic.tower_top_of_injective (hinj : Function.Injective (algebraMap R S)) +theorem Algebra.IsAlgebraic.extendScalars (hinj : Function.Injective (algebraMap R S)) [Algebra.IsAlgebraic R A] : Algebra.IsAlgebraic S A := - ⟨fun _ ↦ _root_.IsAlgebraic.tower_top_of_injective hinj (Algebra.IsAlgebraic.isAlgebraic _)⟩ + ⟨fun _ ↦ _root_.IsAlgebraic.extendScalars hinj (Algebra.IsAlgebraic.isAlgebraic _)⟩ + +@[deprecated (since := "2024-11-18")] +alias Algebra.IsAlgebraic.tower_top_of_injective := Algebra.IsAlgebraic.extendScalars theorem Algebra.IsAlgebraic.tower_bot_of_injective [Algebra.IsAlgebraic R A] (hinj : Function.Injective (algebraMap S A)) : @@ -553,7 +559,7 @@ variable (L) @[stacks 09GF "part one"] theorem IsAlgebraic.tower_top {x : A} (A_alg : IsAlgebraic K x) : IsAlgebraic L x := - A_alg.tower_top_of_injective (algebraMap K L).injective + A_alg.extendScalars (algebraMap K L).injective variable {L} (K) in /-- If `x` is transcendental over `L`, then `x` is transcendental over `K` when @@ -564,7 +570,7 @@ theorem Transcendental.of_tower_top {x : A} (h : Transcendental L x) : /-- If A is an algebraic algebra over K, then A is algebraic over L when L is an extension of K -/ @[stacks 09GF "part two"] theorem Algebra.IsAlgebraic.tower_top [Algebra.IsAlgebraic K A] : Algebra.IsAlgebraic L A := - Algebra.IsAlgebraic.tower_top_of_injective (algebraMap K L).injective + Algebra.IsAlgebraic.extendScalars (algebraMap K L).injective variable (K) @@ -928,7 +934,7 @@ theorem transcendental_supported_polynomial_aeval_X_iff exact Algebra.adjoin_mono (Set.singleton_subset_iff.2 (Set.mem_image_of_mem _ h)) (Polynomial.aeval_mem_adjoin_singleton _ _) · rw [← transcendental_polynomial_aeval_X_iff R i] - refine h.of_tower_top_of_injective fun _ _ heq ↦ MvPolynomial.C_injective σ R ?_ + refine h.restrictScalars fun _ _ heq ↦ MvPolynomial.C_injective σ R ?_ simp_rw [← MvPolynomial.algebraMap_eq] exact congr($(heq).1) diff --git a/Mathlib/RingTheory/AlgebraicIndependent.lean b/Mathlib/RingTheory/AlgebraicIndependent.lean index 6865dee085405..7109a5bc4b402 100644 --- a/Mathlib/RingTheory/AlgebraicIndependent.lean +++ b/Mathlib/RingTheory/AlgebraicIndependent.lean @@ -724,8 +724,7 @@ theorem IsTranscendenceBasis.isAlgebraic_field {F E : Type*} {x : ι → E} (Subalgebra.inclusion (IntermediateField.algebra_adjoin_le_adjoin F S)).toRingHom.toAlgebra haveI : IsScalarTower (adjoin F S) (IntermediateField.adjoin F S) E := IsScalarTower.of_algebraMap_eq (congrFun rfl) - exact Algebra.IsAlgebraic.tower_top_of_injective (R := adjoin F S) - (Subalgebra.inclusion_injective _) + exact Algebra.IsAlgebraic.extendScalars (R := adjoin F S) (Subalgebra.inclusion_injective _) section Field diff --git a/Mathlib/RingTheory/Localization/Integral.lean b/Mathlib/RingTheory/Localization/Integral.lean index d2b5cef36d5e2..07c55c2708e4b 100644 --- a/Mathlib/RingTheory/Localization/Integral.lean +++ b/Mathlib/RingTheory/Localization/Integral.lean @@ -366,7 +366,7 @@ theorem isAlgebraic_iff' [Field K] [IsDomain R] [IsDomain S] [Algebra R K] [Alge rw [div_eq_mul_inv] refine IsIntegral.mul ?_ ?_ · rw [← isAlgebraic_iff_isIntegral] - refine .tower_top_of_injective + refine .extendScalars (NoZeroSMulDivisors.algebraMap_injective R (FractionRing R)) ?_ exact .algebraMap (h a) · rw [← isAlgebraic_iff_isIntegral] From 8d7c6defa0cb5ee614f419923bcc14dbc2f66583 Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Sat, 23 Nov 2024 04:58:47 +0000 Subject: [PATCH 217/829] feat: add `(Algebra|IntermediateField).lift_cardinalMk_adjoin_le` (#19327) Also add `FreeAlgebra.cardinalMk_eq_max_lift`, etc. parallel to `MvPolynomial`. Also fix the name `cardinal_lift_mk_le_max` -> `cardinalMk_le_max_lift` for `MvPolynomial`. --- Mathlib.lean | 1 + Mathlib/Algebra/FreeAlgebra/Cardinality.lean | 78 ++++++++++++++++++++ Mathlib/Algebra/MvPolynomial/Cardinal.lean | 11 ++- Mathlib/FieldTheory/Adjoin.lean | 26 ++++++- 4 files changed, 113 insertions(+), 3 deletions(-) create mode 100644 Mathlib/Algebra/FreeAlgebra/Cardinality.lean diff --git a/Mathlib.lean b/Mathlib.lean index fc3aa1d9991a6..d142e08e33749 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -223,6 +223,7 @@ import Mathlib.Algebra.Field.Subfield.Defs import Mathlib.Algebra.Field.ULift import Mathlib.Algebra.Free import Mathlib.Algebra.FreeAlgebra +import Mathlib.Algebra.FreeAlgebra.Cardinality import Mathlib.Algebra.FreeMonoid.Basic import Mathlib.Algebra.FreeMonoid.Count import Mathlib.Algebra.FreeMonoid.Symbols diff --git a/Mathlib/Algebra/FreeAlgebra/Cardinality.lean b/Mathlib/Algebra/FreeAlgebra/Cardinality.lean new file mode 100644 index 0000000000000..801dea8a9298c --- /dev/null +++ b/Mathlib/Algebra/FreeAlgebra/Cardinality.lean @@ -0,0 +1,78 @@ +/- +Copyright (c) 2024 Jz Pan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jz Pan +-/ +import Mathlib.Algebra.FreeAlgebra +import Mathlib.SetTheory.Cardinal.Free + +/-! +# Cardinality of free algebras + +This file contains some results about the cardinality of `FreeAlgebra`, +parallel to that of `MvPolynomial`. +-/ + +universe u v + +variable (R : Type u) [CommSemiring R] + +open Cardinal + +namespace FreeAlgebra + +variable (X : Type v) + +@[simp] +theorem cardinalMk_eq_max_lift [Nonempty X] [Nontrivial R] : + #(FreeAlgebra R X) = Cardinal.lift.{v} #R ⊔ Cardinal.lift.{u} #X ⊔ ℵ₀ := by + have hX := mk_freeMonoid X + haveI : Infinite (FreeMonoid X) := infinite_iff.2 (by simp [hX]) + rw [equivMonoidAlgebraFreeMonoid.toEquiv.cardinal_eq, MonoidAlgebra, + mk_finsupp_lift_of_infinite, hX, lift_max, lift_aleph0, sup_comm, ← sup_assoc] + +@[simp] +theorem cardinalMk_eq_lift [IsEmpty X] : #(FreeAlgebra R X) = Cardinal.lift.{v} #R := by + have := lift_mk_eq'.2 ⟨show (FreeMonoid X →₀ R) ≃ R from Equiv.finsuppUnique⟩ + rw [lift_id'.{u, v}, lift_umax] at this + rwa [equivMonoidAlgebraFreeMonoid.toEquiv.cardinal_eq, MonoidAlgebra] + +@[nontriviality] +theorem cardinalMk_eq_one [Subsingleton R] : #(FreeAlgebra R X) = 1 := by + rw [equivMonoidAlgebraFreeMonoid.toEquiv.cardinal_eq, MonoidAlgebra, mk_eq_one] + +theorem cardinalMk_le_max_lift : + #(FreeAlgebra R X) ≤ Cardinal.lift.{v} #R ⊔ Cardinal.lift.{u} #X ⊔ ℵ₀ := by + cases subsingleton_or_nontrivial R + · exact (cardinalMk_eq_one R X).trans_le (le_max_of_le_right one_le_aleph0) + cases isEmpty_or_nonempty X + · exact (cardinalMk_eq_lift R X).trans_le (le_max_of_le_left <| le_max_left _ _) + · exact (cardinalMk_eq_max_lift R X).le + +variable (X : Type u) + +theorem cardinalMk_eq_max [Nonempty X] [Nontrivial R] : #(FreeAlgebra R X) = #R ⊔ #X ⊔ ℵ₀ := by + simp + +theorem cardinalMk_eq [IsEmpty X] : #(FreeAlgebra R X) = #R := by + simp + +theorem cardinalMk_le_max : #(FreeAlgebra R X) ≤ #R ⊔ #X ⊔ ℵ₀ := by + simpa using cardinalMk_le_max_lift R X + +end FreeAlgebra + +namespace Algebra + +theorem lift_cardinalMk_adjoin_le {A : Type v} [Semiring A] [Algebra R A] (s : Set A) : + lift.{u} #(adjoin R s) ≤ lift.{v} #R ⊔ lift.{u} #s ⊔ ℵ₀ := by + have H := mk_range_le_lift (f := FreeAlgebra.lift R ((↑) : s → A)) + rw [lift_umax, lift_id'.{v, u}] at H + rw [Algebra.adjoin_eq_range_freeAlgebra_lift] + exact H.trans (FreeAlgebra.cardinalMk_le_max_lift R s) + +theorem cardinalMk_adjoin_le {A : Type u} [Semiring A] [Algebra R A] (s : Set A) : + #(adjoin R s) ≤ #R ⊔ #s ⊔ ℵ₀ := by + simpa using lift_cardinalMk_adjoin_le R s + +end Algebra diff --git a/Mathlib/Algebra/MvPolynomial/Cardinal.lean b/Mathlib/Algebra/MvPolynomial/Cardinal.lean index 3ee977de57b85..09b4c9e178b8c 100644 --- a/Mathlib/Algebra/MvPolynomial/Cardinal.lean +++ b/Mathlib/Algebra/MvPolynomial/Cardinal.lean @@ -40,7 +40,10 @@ theorem cardinalMk_eq_lift [IsEmpty σ] : #(MvPolynomial σ R) = Cardinal.lift.{ @[deprecated (since := "2024-11-10")] alias cardinal_mk_eq_lift := cardinalMk_eq_lift -theorem cardinal_lift_mk_le_max {σ : Type u} {R : Type v} [CommSemiring R] : #(MvPolynomial σ R) ≤ +@[nontriviality] +theorem cardinalMk_eq_one [Subsingleton R] : #(MvPolynomial σ R) = 1 := mk_eq_one _ + +theorem cardinalMk_le_max_lift {σ : Type u} {R : Type v} [CommSemiring R] : #(MvPolynomial σ R) ≤ max (max (Cardinal.lift.{u} #R) <| Cardinal.lift.{v} #σ) ℵ₀ := by cases subsingleton_or_nontrivial R · exact (mk_eq_one _).trans_le (le_max_of_le_right one_le_aleph0) @@ -48,6 +51,8 @@ theorem cardinal_lift_mk_le_max {σ : Type u} {R : Type v} [CommSemiring R] : #( · exact cardinalMk_eq_lift.trans_le (le_max_of_le_left <| le_max_left _ _) · exact cardinalMk_eq_max_lift.le +@[deprecated (since := "2024-11-21")] alias cardinal_lift_mk_le_max := cardinalMk_le_max_lift + end TwoUniverses variable {σ R : Type u} [CommSemiring R] @@ -57,10 +62,12 @@ theorem cardinalMk_eq_max [Nonempty σ] [Nontrivial R] : @[deprecated (since := "2024-11-10")] alias cardinal_mk_eq_max := cardinalMk_eq_max +theorem cardinalMk_eq [IsEmpty σ] : #(MvPolynomial σ R) = #R := by simp + /-- The cardinality of the multivariate polynomial ring, `MvPolynomial σ R` is at most the maximum of `#R`, `#σ` and `ℵ₀` -/ theorem cardinalMk_le_max : #(MvPolynomial σ R) ≤ max (max #R #σ) ℵ₀ := - cardinal_lift_mk_le_max.trans <| by rw [lift_id, lift_id] + cardinalMk_le_max_lift.trans <| by rw [lift_id, lift_id] @[deprecated (since := "2024-11-10")] alias cardinal_mk_le_max := cardinalMk_le_max diff --git a/Mathlib/FieldTheory/Adjoin.lean b/Mathlib/FieldTheory/Adjoin.lean index 9ccf8c539d1f9..51d8f9c940bb7 100644 --- a/Mathlib/FieldTheory/Adjoin.lean +++ b/Mathlib/FieldTheory/Adjoin.lean @@ -11,6 +11,7 @@ import Mathlib.LinearAlgebra.Dimension.FreeAndStrongRankCondition import Mathlib.RingTheory.Adjoin.Dimension import Mathlib.RingTheory.Finiteness.TensorProduct import Mathlib.RingTheory.TensorProduct.Basic +import Mathlib.SetTheory.Cardinal.Subfield /-! # Adjoining Elements to Fields @@ -1680,4 +1681,27 @@ theorem liftAlgHom_fieldRange_eq_of_range_eq (hg : Function.Injective g) end IsFractionRing -set_option linter.style.longFile 1700 +namespace IntermediateField + +universe u v + +open Cardinal + +variable (F : Type u) [Field F] + +theorem lift_cardinalMk_adjoin_le {E : Type v} [Field E] [Algebra F E] (s : Set E) : + Cardinal.lift.{u} #(adjoin F s) ≤ Cardinal.lift.{v} #F ⊔ Cardinal.lift.{u} #s ⊔ ℵ₀ := by + rw [show ↥(adjoin F s) = (adjoin F s).toSubfield from rfl, adjoin_toSubfield] + apply (Cardinal.lift_le.mpr (Subfield.cardinalMk_closure_le_max _)).trans + rw [lift_max, sup_le_iff, lift_aleph0] + refine ⟨(Cardinal.lift_le.mpr ((mk_union_le _ _).trans <| add_le_max _ _)).trans ?_, le_sup_right⟩ + simp_rw [lift_max, lift_aleph0, sup_assoc] + exact sup_le_sup_right mk_range_le_lift _ + +theorem cardinalMk_adjoin_le {E : Type u} [Field E] [Algebra F E] (s : Set E) : + #(adjoin F s) ≤ #F ⊔ #s ⊔ ℵ₀ := by + simpa using lift_cardinalMk_adjoin_le F s + +end IntermediateField + +set_option linter.style.longFile 1800 From a7f12adce3dd754dccc68fb2c83ecb76842f81f0 Mon Sep 17 00:00:00 2001 From: blizzard_inc Date: Sat, 23 Nov 2024 05:51:58 +0000 Subject: [PATCH 218/829] chore(Order/FixedPoints): remove adaptation notes for OrderHom regarding lean4 issue #1910 (#19380) this PR removes 3 porting notes regarding lean4 issue leanprover/lean4#1910, and makes adaptations to mathlib to make use of the newly working notation. Concretely, that means that `OrderHom.lfp f` now is written `f.lfp`, and similar for `OrderHom.dual` and `OrderHom.gfp`. there might still be some `dual x` application left inside code with `open OrderHom`, but those are hard to filter for with grep, as there are many declarations named `dual` in different namespaces. --- Mathlib/Order/FixedPoints.lean | 98 +++++++++---------- Mathlib/Order/Hom/Basic.lean | 7 +- Mathlib/Order/Hom/Bounded.lean | 6 +- Mathlib/Order/Interval/Basic.lean | 4 +- .../Cardinal/SchroederBernstein.lean | 3 +- .../Ordinal/FixedPointApproximants.lean | 32 +++--- 6 files changed, 71 insertions(+), 79 deletions(-) diff --git a/Mathlib/Order/FixedPoints.lean b/Mathlib/Order/FixedPoints.lean index e6f40735b5ddc..20419434bbcf6 100644 --- a/Mathlib/Order/FixedPoints.lean +++ b/Mathlib/Order/FixedPoints.lean @@ -51,76 +51,73 @@ def gfp : (α →o α) →o α where toFun f := sSup { a | a ≤ f a } monotone' _ _ hle := sSup_le_sSup fun a ha => le_trans ha (hle a) -theorem lfp_le {a : α} (h : f a ≤ a) : lfp f ≤ a := +theorem lfp_le {a : α} (h : f a ≤ a) : f.lfp ≤ a := sInf_le h -theorem lfp_le_fixed {a : α} (h : f a = a) : lfp f ≤ a := +theorem lfp_le_fixed {a : α} (h : f a = a) : f.lfp ≤ a := f.lfp_le h.le -theorem le_lfp {a : α} (h : ∀ b, f b ≤ b → a ≤ b) : a ≤ lfp f := +theorem le_lfp {a : α} (h : ∀ b, f b ≤ b → a ≤ b) : a ≤ f.lfp := le_sInf h --- Porting note: for the rest of the file, replace the dot notation `_.lfp` with `lfp _` --- same for `_.gfp`, `_.dual` --- Probably related to https://github.com/leanprover/lean4/issues/1910 -theorem map_le_lfp {a : α} (ha : a ≤ lfp f) : f a ≤ lfp f := +theorem map_le_lfp {a : α} (ha : a ≤ f.lfp) : f a ≤ f.lfp := f.le_lfp fun _ hb => (f.mono <| le_sInf_iff.1 ha _ hb).trans hb @[simp] -theorem map_lfp : f (lfp f) = lfp f := - have h : f (lfp f) ≤ lfp f := f.map_le_lfp le_rfl +theorem map_lfp : f f.lfp = f.lfp := + have h : f f.lfp ≤ f.lfp := f.map_le_lfp le_rfl h.antisymm <| f.lfp_le <| f.mono h -theorem isFixedPt_lfp : IsFixedPt f (lfp f) := +theorem isFixedPt_lfp : IsFixedPt f f.lfp := f.map_lfp -theorem lfp_le_map {a : α} (ha : lfp f ≤ a) : lfp f ≤ f a := +theorem lfp_le_map {a : α} (ha : f.lfp ≤ a) : f.lfp ≤ f a := calc - lfp f = f (lfp f) := f.map_lfp.symm + f.lfp = f f.lfp := f.map_lfp.symm _ ≤ f a := f.mono ha -theorem isLeast_lfp_le : IsLeast { a | f a ≤ a } (lfp f) := +theorem isLeast_lfp_le : IsLeast { a | f a ≤ a } f.lfp := ⟨f.map_lfp.le, fun _ => f.lfp_le⟩ -theorem isLeast_lfp : IsLeast (fixedPoints f) (lfp f) := +theorem isLeast_lfp : IsLeast (fixedPoints f) f.lfp := ⟨f.isFixedPt_lfp, fun _ => f.lfp_le_fixed⟩ -theorem lfp_induction {p : α → Prop} (step : ∀ a, p a → a ≤ lfp f → p (f a)) - (hSup : ∀ s, (∀ a ∈ s, p a) → p (sSup s)) : p (lfp f) := by - set s := { a | a ≤ lfp f ∧ p a } +theorem lfp_induction {p : α → Prop} (step : ∀ a, p a → a ≤ f.lfp → p (f a)) + (hSup : ∀ s, (∀ a ∈ s, p a) → p (sSup s)) : p f.lfp := by + set s := { a | a ≤ f.lfp ∧ p a } specialize hSup s fun a => And.right - suffices sSup s = lfp f from this ▸ hSup - have h : sSup s ≤ lfp f := sSup_le fun b => And.left + suffices sSup s = f.lfp from this ▸ hSup + have h : sSup s ≤ f.lfp := sSup_le fun b => And.left have hmem : f (sSup s) ∈ s := ⟨f.map_le_lfp h, step _ hSup h⟩ exact h.antisymm (f.lfp_le <| le_sSup hmem) -theorem le_gfp {a : α} (h : a ≤ f a) : a ≤ gfp f := +theorem le_gfp {a : α} (h : a ≤ f a) : a ≤ f.gfp := le_sSup h -theorem gfp_le {a : α} (h : ∀ b, b ≤ f b → b ≤ a) : gfp f ≤ a := +theorem gfp_le {a : α} (h : ∀ b, b ≤ f b → b ≤ a) : f.gfp ≤ a := sSup_le h -theorem isFixedPt_gfp : IsFixedPt f (gfp f) := +theorem isFixedPt_gfp : IsFixedPt f f.gfp := f.dual.isFixedPt_lfp @[simp] -theorem map_gfp : f (gfp f) = gfp f := +theorem map_gfp : f f.gfp = f.gfp := f.dual.map_lfp -theorem map_le_gfp {a : α} (ha : a ≤ gfp f) : f a ≤ gfp f := +theorem map_le_gfp {a : α} (ha : a ≤ f.gfp) : f a ≤ f.gfp := f.dual.lfp_le_map ha -theorem gfp_le_map {a : α} (ha : gfp f ≤ a) : gfp f ≤ f a := +theorem gfp_le_map {a : α} (ha : f.gfp ≤ a) : f.gfp ≤ f a := f.dual.map_le_lfp ha -theorem isGreatest_gfp_le : IsGreatest { a | a ≤ f a } (gfp f) := +theorem isGreatest_gfp_le : IsGreatest { a | a ≤ f a } f.gfp := f.dual.isLeast_lfp_le -theorem isGreatest_gfp : IsGreatest (fixedPoints f) (gfp f) := +theorem isGreatest_gfp : IsGreatest (fixedPoints f) f.gfp := f.dual.isLeast_lfp -theorem gfp_induction {p : α → Prop} (step : ∀ a, p a → gfp f ≤ a → p (f a)) - (hInf : ∀ s, (∀ a ∈ s, p a) → p (sInf s)) : p (gfp f) := +theorem gfp_induction {p : α → Prop} (step : ∀ a, p a → f.gfp ≤ a → p (f a)) + (hInf : ∀ s, (∀ a ∈ s, p a) → p (sInf s)) : p f.gfp := f.dual.lfp_induction step hInf end Basic @@ -130,27 +127,26 @@ section Eqn variable [CompleteLattice α] [CompleteLattice β] (f : β →o α) (g : α →o β) -- Rolling rule -theorem map_lfp_comp : f (lfp (g.comp f)) = lfp (f.comp g) := +theorem map_lfp_comp : f (g.comp f).lfp = (f.comp g).lfp := le_antisymm ((f.comp g).map_lfp ▸ f.mono (lfp_le_fixed _ <| congr_arg g (f.comp g).map_lfp)) <| lfp_le _ (congr_arg f (g.comp f).map_lfp).le -theorem map_gfp_comp : f (gfp (g.comp f)) = gfp (f.comp g) := - f.dual.map_lfp_comp (OrderHom.dual g) +theorem map_gfp_comp : f (g.comp f).gfp = (f.comp g).gfp := + f.dual.map_lfp_comp g.dual -- Diagonal rule -theorem lfp_lfp (h : α →o α →o α) : lfp (lfp.comp h) = lfp h.onDiag := by - let a := lfp (lfp.comp h) +theorem lfp_lfp (h : α →o α →o α) : (lfp.comp h).lfp = h.onDiag.lfp := by + let a := (lfp.comp h).lfp refine (lfp_le _ ?_).antisymm (lfp_le _ (Eq.le ?_)) · exact lfp_le _ h.onDiag.map_lfp.le have ha : (lfp ∘ h) a = a := (lfp.comp h).map_lfp calc - h a a = h a (lfp (h a)) := congr_arg (h a) ha.symm - _ = lfp (h a) := (h a).map_lfp + h a a = h a (h a).lfp := congr_arg (h a) ha.symm + _ = (h a).lfp := (h a).map_lfp _ = a := ha -theorem gfp_gfp (h : α →o α →o α) : gfp (gfp.comp h) = gfp h.onDiag := - @lfp_lfp αᵒᵈ _ <| (OrderHom.dualIso αᵒᵈ αᵒᵈ).symm.toOrderEmbedding.toOrderHom.comp - (OrderHom.dual h) +theorem gfp_gfp (h : α →o α →o α) : (gfp.comp h).gfp = h.onDiag.gfp := + @lfp_lfp αᵒᵈ _ <| (OrderHom.dualIso αᵒᵈ αᵒᵈ).symm.toOrderEmbedding.toOrderHom.comp h.dual end Eqn @@ -158,25 +154,25 @@ section PrevNext variable [CompleteLattice α] (f : α →o α) -theorem gfp_const_inf_le (x : α) : gfp (const α x ⊓ f) ≤ x := +theorem gfp_const_inf_le (x : α) : (const α x ⊓ f).gfp ≤ x := (gfp_le _) fun _ hb => hb.trans inf_le_left /-- Previous fixed point of a monotone map. If `f` is a monotone self-map of a complete lattice and `x` is a point such that `f x ≤ x`, then `f.prevFixed x hx` is the greatest fixed point of `f` that is less than or equal to `x`. -/ def prevFixed (x : α) (hx : f x ≤ x) : fixedPoints f := - ⟨gfp (const α x ⊓ f), + ⟨(const α x ⊓ f).gfp, calc - f (gfp (const α x ⊓ f)) = x ⊓ f (gfp (const α x ⊓ f)) := + f (const α x ⊓ f).gfp = x ⊓ f (const α x ⊓ f).gfp := Eq.symm <| inf_of_le_right <| (f.mono <| f.gfp_const_inf_le x).trans hx - _ = gfp (const α x ⊓ f) := (const α x ⊓ f).map_gfp + _ = (const α x ⊓ f).gfp := (const α x ⊓ f).map_gfp ⟩ /-- Next fixed point of a monotone map. If `f` is a monotone self-map of a complete lattice and `x` is a point such that `x ≤ f x`, then `f.nextFixed x hx` is the least fixed point of `f` that is greater than or equal to `x`. -/ def nextFixed (x : α) (hx : x ≤ f x) : fixedPoints f := - { f.dual.prevFixed x hx with val := lfp (const α x ⊔ f) } + { f.dual.prevFixed x hx with val := (const α x ⊔ f).lfp } theorem prevFixed_le {x : α} (hx : f x ≤ x) : ↑(f.prevFixed x hx) ≤ x := f.gfp_const_inf_le x @@ -240,7 +236,7 @@ instance : SemilatticeSup (fixedPoints f) := /- porting note: removed `Subtype.partialOrder _` from mathlib3port version, threw `typeclass instance` error and was seemingly unnecessary?-/ instance : SemilatticeInf (fixedPoints f) := - { OrderDual.instSemilatticeInf (fixedPoints (OrderHom.dual f)) with + { OrderDual.instSemilatticeInf (fixedPoints f.dual) with inf := fun x y => f.prevFixed (x ⊓ y) (f.map_inf_fixedPoints_le x y) } -- Porting note: `coe` replaced with `Subtype.val` @@ -271,8 +267,8 @@ instance completeLattice : CompleteLattice (fixedPoints f) where __ := inferInstanceAs (SemilatticeSup (fixedPoints f)) __ := inferInstanceAs (CompleteSemilatticeInf (fixedPoints f)) __ := inferInstanceAs (CompleteSemilatticeSup (fixedPoints f)) - top := ⟨gfp f, f.isFixedPt_gfp⟩ - bot := ⟨lfp f, f.isFixedPt_lfp⟩ + top := ⟨f.gfp, f.isFixedPt_gfp⟩ + bot := ⟨f.lfp, f.isFixedPt_lfp⟩ le_top x := f.le_gfp x.2.ge bot_le x := f.lfp_le x.2.le @@ -281,7 +277,7 @@ open OmegaCompletePartialOrder fixedPoints /-- **Kleene's fixed point Theorem**: The least fixed point in a complete lattice is the supremum of iterating a function on bottom arbitrary often. -/ theorem lfp_eq_sSup_iterate (h : ωScottContinuous f) : - lfp f = ⨆ n, f^[n] ⊥ := by + f.lfp = ⨆ n, f^[n] ⊥ := by apply le_antisymm · apply lfp_le_fixed exact Function.mem_fixedPoints.mp (ωSup_iterate_mem_fixedPoint @@ -290,8 +286,8 @@ theorem lfp_eq_sSup_iterate (h : ωScottContinuous f) : intro a h_a exact ωSup_iterate_le_prefixedPoint ⟨f, h.map_ωSup_of_orderHom⟩ ⊥ bot_le h_a bot_le -theorem gfp_eq_sInf_iterate (h : ωScottContinuous (OrderHom.dual f)) : - gfp f = ⨅ n, f^[n] ⊤ := - lfp_eq_sSup_iterate (OrderHom.dual f) h +theorem gfp_eq_sInf_iterate (h : ωScottContinuous f.dual) : + f.gfp = ⨅ n, f^[n] ⊤ := + lfp_eq_sSup_iterate f.dual h end fixedPoints diff --git a/Mathlib/Order/Hom/Basic.lean b/Mathlib/Order/Hom/Basic.lean index aaa22efefcb3e..ef9ebbb809b7f 100644 --- a/Mathlib/Order/Hom/Basic.lean +++ b/Mathlib/Order/Hom/Basic.lean @@ -488,16 +488,13 @@ protected def dual : (α →o β) ≃ (αᵒᵈ →o βᵒᵈ) where left_inv _ := rfl right_inv _ := rfl --- Porting note: We used to be able to write `(OrderHom.id : α →o α).dual` here rather than --- `OrderHom.dual (OrderHom.id : α →o α)`. --- See https://github.com/leanprover/lean4/issues/1910 @[simp] -theorem dual_id : OrderHom.dual (OrderHom.id : α →o α) = OrderHom.id := +theorem dual_id : (OrderHom.id : α →o α).dual = OrderHom.id := rfl @[simp] theorem dual_comp (g : β →o γ) (f : α →o β) : - OrderHom.dual (g.comp f) = (OrderHom.dual g).comp (OrderHom.dual f) := + (g.comp f).dual = g.dual.comp f.dual := rfl @[simp] diff --git a/Mathlib/Order/Hom/Bounded.lean b/Mathlib/Order/Hom/Bounded.lean index 00886351437d4..6466c90ddaf43 100644 --- a/Mathlib/Order/Hom/Bounded.lean +++ b/Mathlib/Order/Hom/Bounded.lean @@ -725,18 +725,18 @@ protected def dual : BoundedOrderHom α β ≃ BoundedOrderHom αᵒᵈ βᵒᵈ where - toFun f := ⟨OrderHom.dual f.toOrderHom, f.map_bot', f.map_top'⟩ + toFun f := ⟨f.toOrderHom.dual, f.map_bot', f.map_top'⟩ invFun f := ⟨OrderHom.dual.symm f.toOrderHom, f.map_bot', f.map_top'⟩ left_inv _ := ext fun _ => rfl right_inv _ := ext fun _ => rfl @[simp] -theorem dual_id : BoundedOrderHom.dual (BoundedOrderHom.id α) = BoundedOrderHom.id _ := +theorem dual_id : (BoundedOrderHom.id α).dual = BoundedOrderHom.id _ := rfl @[simp] theorem dual_comp (g : BoundedOrderHom β γ) (f : BoundedOrderHom α β) : - BoundedOrderHom.dual (g.comp f) = g.dual.comp (BoundedOrderHom.dual f) := + (g.comp f).dual = g.dual.comp f.dual := rfl @[simp] diff --git a/Mathlib/Order/Interval/Basic.lean b/Mathlib/Order/Interval/Basic.lean index 9ef57523db1f8..014a33ce871f3 100644 --- a/Mathlib/Order/Interval/Basic.lean +++ b/Mathlib/Order/Interval/Basic.lean @@ -153,7 +153,7 @@ theorem map_map (g : β →o γ) (f : α →o β) (a : NonemptyInterval α) : @[simp] theorem dual_map (f : α →o β) (a : NonemptyInterval α) : - dual (a.map f) = a.dual.map (OrderHom.dual f) := + dual (a.map f) = a.dual.map f.dual := rfl /-- Binary pushforward of nonempty intervals. -/ @@ -360,7 +360,7 @@ theorem map_map (g : β →o γ) (f : α →o β) (s : Interval α) : (s.map f). Option.map_map _ _ _ @[simp] -theorem dual_map (f : α →o β) (s : Interval α) : dual (s.map f) = s.dual.map (OrderHom.dual f) := by +theorem dual_map (f : α →o β) (s : Interval α) : dual (s.map f) = s.dual.map f.dual := by cases s · rfl · exact WithBot.map_comm rfl _ diff --git a/Mathlib/SetTheory/Cardinal/SchroederBernstein.lean b/Mathlib/SetTheory/Cardinal/SchroederBernstein.lean index a5c9f5c018629..02acdbd1461b6 100644 --- a/Mathlib/SetTheory/Cardinal/SchroederBernstein.lean +++ b/Mathlib/SetTheory/Cardinal/SchroederBernstein.lean @@ -49,8 +49,7 @@ theorem schroeder_bernstein {f : α → β} {g : β → α} (hf : Function.Injec { toFun := fun s => (g '' (f '' s)ᶜ)ᶜ monotone' := fun s t hst => compl_subset_compl.mpr <| image_subset _ <| compl_subset_compl.mpr <| image_subset _ hst } - -- Porting note: dot notation `F.lfp` doesn't work here - set s : Set α := OrderHom.lfp F + set s : Set α := F.lfp have hs : (g '' (f '' s)ᶜ)ᶜ = s := F.map_lfp have hns : g '' (f '' s)ᶜ = sᶜ := compl_injective (by simp [hs]) set g' := invFun g diff --git a/Mathlib/SetTheory/Ordinal/FixedPointApproximants.lean b/Mathlib/SetTheory/Ordinal/FixedPointApproximants.lean index 97272de665f4f..4d1e937f10243 100644 --- a/Mathlib/SetTheory/Ordinal/FixedPointApproximants.lean +++ b/Mathlib/SetTheory/Ordinal/FixedPointApproximants.lean @@ -226,9 +226,9 @@ theorem lfpApprox_le_of_mem_fixedPoints {a : α} /-- The approximation sequence converges at the successor of the domain's cardinality to the least fixed point if starting from `⊥` -/ -theorem lfpApprox_ord_eq_lfp : lfpApprox f ⊥ (ord <| succ #α) = lfp f := by +theorem lfpApprox_ord_eq_lfp : lfpApprox f ⊥ (ord <| succ #α) = f.lfp := by apply le_antisymm - · have h_lfp : ∃ y : fixedPoints f, lfp f = y := by use ⊥; exact rfl + · have h_lfp : ∃ y : fixedPoints f, f.lfp = y := by use ⊥; exact rfl let ⟨y, h_y⟩ := h_lfp; rw [h_y] exact lfpApprox_le_of_mem_fixedPoints f ⊥ y.2 bot_le (ord <| succ #α) · have h_fix : ∃ y : fixedPoints f, lfpApprox f ⊥ (ord <| succ #α) = y := by @@ -238,7 +238,7 @@ theorem lfpApprox_ord_eq_lfp : lfpApprox f ⊥ (ord <| succ #α) = lfp f := by exact lfp_le_fixed f x.prop /-- Some approximation of the least fixed point starting from `⊥` is the least fixed point. -/ -theorem lfp_mem_range_lfpApprox : lfp f ∈ Set.range (lfpApprox f ⊥) := by +theorem lfp_mem_range_lfpApprox : f.lfp ∈ Set.range (lfpApprox f ⊥) := by use ord <| succ #α exact lfpApprox_ord_eq_lfp f @@ -256,52 +256,52 @@ decreasing_by exact h unseal gfpApprox lfpApprox theorem gfpApprox_antitone : Antitone (gfpApprox f x) := - lfpApprox_monotone (OrderHom.dual f) x + lfpApprox_monotone f.dual x theorem gfpApprox_le {a : Ordinal} : gfpApprox f x a ≤ x := - le_lfpApprox (OrderHom.dual f) x + le_lfpApprox f.dual x theorem gfpApprox_add_one (h : f x ≤ x) (a : Ordinal) : gfpApprox f x (a+1) = f (gfpApprox f x a) := - lfpApprox_add_one (OrderHom.dual f) x h a + lfpApprox_add_one f.dual x h a theorem gfpApprox_mono_left : Monotone (gfpApprox : (α →o α) → _) := by intro f g h - have : OrderHom.dual g ≤ OrderHom.dual f := h + have : g.dual ≤ f.dual := h exact lfpApprox_mono_left this theorem gfpApprox_mono_mid : Monotone (gfpApprox f) := - fun _ _ h => lfpApprox_mono_mid (OrderHom.dual f) h + fun _ _ h => lfpApprox_mono_mid f.dual h /-- The approximations of the greatest fixed point stabilize at a fixed point of `f` -/ theorem gfpApprox_eq_of_mem_fixedPoints {a b : Ordinal} (h_init : f x ≤ x) (h_ab : a ≤ b) (h : gfpApprox f x a ∈ fixedPoints f) : gfpApprox f x b = gfpApprox f x a := - lfpApprox_eq_of_mem_fixedPoints (OrderHom.dual f) x h_init h_ab h + lfpApprox_eq_of_mem_fixedPoints f.dual x h_init h_ab h /-- There are distinct indices smaller than the successor of the domain's cardinality yielding the same value -/ theorem exists_gfpApprox_eq_gfpApprox : ∃ a < ord <| succ #α, ∃ b < ord <| succ #α, a ≠ b ∧ gfpApprox f x a = gfpApprox f x b := - exists_lfpApprox_eq_lfpApprox (OrderHom.dual f) x + exists_lfpApprox_eq_lfpApprox f.dual x /-- The approximation at the index of the successor of the domain's cardinality is a fixed point -/ lemma gfpApprox_ord_mem_fixedPoint (h_init : f x ≤ x) : gfpApprox f x (ord <| succ #α) ∈ fixedPoints f := - lfpApprox_ord_mem_fixedPoint (OrderHom.dual f) x h_init + lfpApprox_ord_mem_fixedPoint f.dual x h_init /-- Every value of the approximation is greater or equal than every fixed point of `f` less or equal than the initial value -/ lemma le_gfpApprox_of_mem_fixedPoints {a : α} (h_a : a ∈ fixedPoints f) (h_le_init : a ≤ x) (i : Ordinal) : a ≤ gfpApprox f x i := - lfpApprox_le_of_mem_fixedPoints (OrderHom.dual f) x h_a h_le_init i + lfpApprox_le_of_mem_fixedPoints f.dual x h_a h_le_init i /-- The approximation sequence converges at the successor of the domain's cardinality to the greatest fixed point if starting from `⊥` -/ -theorem gfpApprox_ord_eq_gfp : gfpApprox f ⊤ (ord <| succ #α) = gfp f := - lfpApprox_ord_eq_lfp (OrderHom.dual f) +theorem gfpApprox_ord_eq_gfp : gfpApprox f ⊤ (ord <| succ #α) = f.gfp := + lfpApprox_ord_eq_lfp f.dual /-- Some approximation of the least fixed point starting from `⊤` is the greatest fixed point. -/ -theorem gfp_mem_range_gfpApprox : gfp f ∈ Set.range (gfpApprox f ⊤) := - lfp_mem_range_lfpApprox (OrderHom.dual f) +theorem gfp_mem_range_gfpApprox : f.gfp ∈ Set.range (gfpApprox f ⊤) := + lfp_mem_range_lfpApprox f.dual end OrdinalApprox From f700f821a8e256bf132e1e9fb6c433f83a5ceda4 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Sat, 23 Nov 2024 06:30:16 +0000 Subject: [PATCH 219/829] feat(FieldTheory): minimal polynomials determine an algebraic extension (Isaacs 1980) (#18412) If E/F is an algebraic extension and K/F is an arbitrary extension, and if the minimal polynomial of every element of E over F has a root in K, we show there exists an F-embedding from E into K. If E/F and K/F have the same set of minimal polynomials, they are then isomorphic. Another corollary is that an algebraic extension E/F is an algebraic closure if every (monic irreducible) polynomial in F[X] has a root in E. These results are proven in FieldTheory/Isaacs. An important definition (`IsExtendible`) and two important lemmas towards these theorems concern `IntermediateField.Lifts` and reside in FieldTheory/Extension. The lemmas say that the supremum of a chain of extendible lifts is also extendible, and that a lift with full domain exists if the bottom lift is extendible. We do slight refactoring to extract `Lifts.union`, and provide new API lemmas for Lifts. Co-authored-by: Junyan Xu --- Mathlib.lean | 1 + Mathlib/FieldTheory/Extension.lean | 162 +++++++++++++++++++-- Mathlib/FieldTheory/IsAlgClosed/Basic.lean | 8 +- Mathlib/FieldTheory/Isaacs.lean | 115 +++++++++++++++ docs/references.bib | 13 ++ 5 files changed, 280 insertions(+), 19 deletions(-) create mode 100644 Mathlib/FieldTheory/Isaacs.lean diff --git a/Mathlib.lean b/Mathlib.lean index d142e08e33749..075b258277b1b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2951,6 +2951,7 @@ import Mathlib.FieldTheory.IsAlgClosed.Classification import Mathlib.FieldTheory.IsAlgClosed.Spectrum import Mathlib.FieldTheory.IsPerfectClosure import Mathlib.FieldTheory.IsSepClosed +import Mathlib.FieldTheory.Isaacs import Mathlib.FieldTheory.JacobsonNoether import Mathlib.FieldTheory.KrullTopology import Mathlib.FieldTheory.KummerExtension diff --git a/Mathlib/FieldTheory/Extension.lean b/Mathlib/FieldTheory/Extension.lean index da3e16176910d..142e5ed48085b 100644 --- a/Mathlib/FieldTheory/Extension.lean +++ b/Mathlib/FieldTheory/Extension.lean @@ -1,8 +1,9 @@ /- Copyright (c) 2020 Thomas Browning. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Thomas Browning +Authors: Thomas Browning, Junyan Xu -/ +import Mathlib.Data.Fintype.Order import Mathlib.FieldTheory.Adjoin /-! @@ -12,6 +13,12 @@ import Mathlib.FieldTheory.Adjoin field extensions, K is another extension of F, and `f` is an embedding of L/F into K/F, such that the minimal polynomials of a set of generators of E/L splits in K (via `f`), then `f` extends to an embedding of E/F into K/F. + +## Reference + +[Isaacs1980] *Roots of Polynomials in Algebraic Extensions of Fields*, +The American Mathematical Monthly + -/ open Polynomial @@ -29,6 +36,8 @@ structure Lifts where variable {F E K} +namespace Lifts + instance : PartialOrder (Lifts F E K) where le L₁ L₂ := ∃ h : L₁.carrier ≤ L₂.carrier, ∀ x, L₂.emb (inclusion h x) = L₁.emb x le_refl L := ⟨le_rfl, by simp⟩ @@ -52,25 +61,144 @@ noncomputable instance : OrderBot (Lifts F E K) where noncomputable instance : Inhabited (Lifts F E K) := ⟨⊥⟩ -/-- A chain of lifts has an upper bound. -/ -theorem Lifts.exists_upper_bound (c : Set (Lifts F E K)) (hc : IsChain (· ≤ ·) c) : - ∃ ub, ∀ a ∈ c, a ≤ ub := by +variable {L₁ L₂ : Lifts F E K} + +theorem le_iff : L₁ ≤ L₂ ↔ + ∃ h : L₁.carrier ≤ L₂.carrier, L₂.emb.comp (inclusion h) = L₁.emb := by + simp_rw [AlgHom.ext_iff]; rfl + +theorem eq_iff_le_carrier_eq : L₁ = L₂ ↔ L₁ ≤ L₂ ∧ L₁.carrier = L₂.carrier := + ⟨fun eq ↦ ⟨eq.le, congr_arg _ eq⟩, fun ⟨le, eq⟩ ↦ le.antisymm ⟨eq.ge, fun x ↦ (le.2 ⟨x, _⟩).symm⟩⟩ + +theorem eq_iff : L₁ = L₂ ↔ + ∃ h : L₁.carrier = L₂.carrier, L₂.emb.comp (inclusion h.le) = L₁.emb := by + rw [eq_iff_le_carrier_eq, le_iff] + exact ⟨fun h ↦ ⟨h.2, h.1.2⟩, fun h ↦ ⟨⟨h.1.le, h.2⟩, h.1⟩⟩ + +theorem lt_iff_le_carrier_ne : L₁ < L₂ ↔ L₁ ≤ L₂ ∧ L₁.carrier ≠ L₂.carrier := by + rw [lt_iff_le_and_ne, and_congr_right]; intro h; simp_rw [Ne, eq_iff_le_carrier_eq, h, true_and] + +theorem lt_iff : L₁ < L₂ ↔ + ∃ h : L₁.carrier < L₂.carrier, L₂.emb.comp (inclusion h.le) = L₁.emb := by + rw [lt_iff_le_carrier_ne, le_iff] + exact ⟨fun h ↦ ⟨h.1.1.lt_of_ne h.2, h.1.2⟩, fun h ↦ ⟨⟨h.1.le, h.2⟩, h.1.ne⟩⟩ + +theorem le_of_carrier_le_iSup {ι} {ρ : ι → Lifts F E K} {σ τ : Lifts F E K} + (hσ : ∀ i, ρ i ≤ σ) (hτ : ∀ i, ρ i ≤ τ) (carrier_le : σ.carrier ≤ ⨆ i, (ρ i).carrier) : + σ ≤ τ := + le_iff.mpr ⟨carrier_le.trans (iSup_le fun i ↦ (hτ i).1), algHom_ext_of_eq_adjoin _ + (carrier_le.antisymm (iSup_le fun i ↦ (hσ i).1)|>.trans <| iSup_eq_adjoin _ _) fun x hx ↦ + have ⟨i, hx⟩ := Set.mem_iUnion.mp hx + ((hτ i).2 ⟨x, hx⟩).trans ((hσ i).2 ⟨x, hx⟩).symm⟩ + +/-- `σ : L →ₐ[F] K` is an extendible lift ("extendible pair" in [Isaacs1980]) if for every +intermediate field `M` that is finite-dimensional over `L`, `σ` extends to some `M →ₐ[F] K`. +In our definition we only require `M` to be finitely generated over `L`, which is equivalent +if the ambient field `E` is algebraic over `F` (which is the case in our main application). +We also allow the domain of the extension to be an intermediate field that properly contains `M`, +since one can always restrict the domain to `M`. -/ +def IsExtendible (σ : Lifts F E K) : Prop := + ∀ S : Finset E, ∃ τ ≥ σ, (S : Set E) ⊆ τ.carrier + +section Chain +variable (c : Set (Lifts F E K)) (hc : IsChain (· ≤ ·) c) + +/-- The union of a chain of lifts. -/ +noncomputable def union : Lifts F E K := let t (i : ↑(insert ⊥ c)) := i.val.carrier - let t' (i) := (t i).toSubalgebra have hc := hc.insert fun _ _ _ ↦ .inl bot_le have dir : Directed (· ≤ ·) t := hc.directedOn.directed_val.mono_comp _ fun _ _ h ↦ h.1 - refine ⟨⟨iSup t, (Subalgebra.iSupLift t' dir (fun i ↦ i.val.emb) (fun i j h ↦ ?_) _ rfl).comp - (Subalgebra.equivOfEq _ _ <| toSubalgebra_iSup_of_directed dir)⟩, - fun L hL ↦ have hL := Set.mem_insert_of_mem ⊥ hL; ⟨le_iSup t ⟨L, hL⟩, fun x ↦ ?_⟩⟩ - · refine AlgHom.ext fun x ↦ (hc.total i.2 j.2).elim (fun hij ↦ (hij.snd x).symm) fun hji ↦ ?_ - erw [AlgHom.comp_apply, ← hji.snd (Subalgebra.inclusion h x), - inclusion_inclusion, inclusion_self, AlgHom.id_apply x] - · dsimp only [AlgHom.comp_apply] - exact Subalgebra.iSupLift_inclusion (K := t') (i := ⟨L, hL⟩) x (le_iSup t' ⟨L, hL⟩) + ⟨iSup t, (Subalgebra.iSupLift (toSubalgebra <| t ·) dir (·.val.emb) (fun i j h ↦ + AlgHom.ext fun x ↦ (hc.total i.2 j.2).elim (fun hij ↦ (hij.snd x).symm) fun hji ↦ by + erw [AlgHom.comp_apply, ← hji.snd (Subalgebra.inclusion h x), + inclusion_inclusion, inclusion_self, AlgHom.id_apply x]) _ rfl).comp + (Subalgebra.equivOfEq _ _ <| toSubalgebra_iSup_of_directed dir)⟩ + +theorem le_union ⦃σ : Lifts F E K⦄ (hσ : σ ∈ c) : σ ≤ union c hc := + have hσ := Set.mem_insert_of_mem ⊥ hσ + let t (i : ↑(insert ⊥ c)) := i.val.carrier + ⟨le_iSup t ⟨σ, hσ⟩, fun x ↦ by + dsimp only [union, AlgHom.comp_apply] + exact Subalgebra.iSupLift_inclusion (K := (toSubalgebra <| t ·)) + (i := ⟨σ, hσ⟩) x (le_iSup (toSubalgebra <| t ·) ⟨σ, hσ⟩)⟩ + +theorem carrier_union : (union c hc).carrier = ⨆ i : c, i.1.carrier := + le_antisymm (iSup_le <| by rintro ⟨i, rfl|hi⟩; exacts [bot_le, le_iSup_of_le ⟨i, hi⟩ le_rfl]) <| + iSup_le fun i ↦ le_iSup_of_le ⟨i, .inr i.2⟩ le_rfl + +/-- A chain of lifts has an upper bound. -/ +theorem exists_upper_bound (c : Set (Lifts F E K)) (hc : IsChain (· ≤ ·) c) : + ∃ ub, ∀ a ∈ c, a ≤ ub := ⟨_, le_union c hc⟩ + +theorem union_isExtendible [alg : Algebra.IsAlgebraic F E] + [Nonempty c] (hext : ∀ σ ∈ c, σ.IsExtendible) : + (union c hc).IsExtendible := fun S ↦ by + let Ω := adjoin F (S : Set E) →ₐ[F] K + have ⟨ω, hω⟩ : ∃ ω : Ω, ∀ π : c, ∃ θ ≥ π.1, ⟨_, ω⟩ ≤ θ ∧ θ.carrier = π.1.1 ⊔ adjoin F S := by + by_contra!; choose π hπ using this + have := finiteDimensional_adjoin (S := (S : Set E)) fun _ _ ↦ (alg.isIntegral).1 _ + have ⟨π₀, hπ₀⟩ := hc.directed.finite_le π + have ⟨θ, hθπ, hθ⟩ := hext _ π₀.2 S + rw [← adjoin_le_iff] at hθ + let θ₀ := θ.emb.comp (inclusion hθ) + have := (hπ₀ θ₀).trans hθπ + exact hπ θ₀ ⟨_, θ.emb.comp <| inclusion <| sup_le this.1 hθ⟩ + ⟨le_sup_left, this.2⟩ ⟨le_sup_right, fun _ ↦ rfl⟩ rfl + choose θ ge hθ eq using hω + have : IsChain (· ≤ ·) (Set.range θ) := by + simp_rw [← restrictScalars_adjoin_eq_sup, restrictScalars_adjoin] at eq + rintro _ ⟨π₁, rfl⟩ _ ⟨π₂, rfl⟩ - + wlog h : π₁ ≤ π₂ generalizing π₁ π₂ + · exact (this _ _ <| (hc.total π₁.2 π₂.2).resolve_left h).symm + refine .inl (le_iff.mpr ⟨?_, algHom_ext_of_eq_adjoin _ (eq _) ?_⟩) + · rw [eq, eq]; exact adjoin.mono _ _ _ (Set.union_subset_union_left _ h.1) + rintro x (hx|hx) + · change (θ π₂).emb (inclusion (ge π₂).1 <| inclusion h.1 ⟨x, hx⟩) = + (θ π₁).emb (inclusion (ge π₁).1 ⟨x, hx⟩) + rw [(ge π₁).2, (ge π₂).2, h.2] + · change (θ π₂).emb (inclusion (hθ π₂).1 ⟨x, subset_adjoin _ _ hx⟩) = + (θ π₁).emb (inclusion (hθ π₁).1 ⟨x, subset_adjoin _ _ hx⟩) + rw [(hθ π₁).2, (hθ π₂).2] + refine ⟨union _ this, le_of_carrier_le_iSup (fun π ↦ le_union c hc π.2) + (fun π ↦ (ge π).trans <| le_union _ _ ⟨_, rfl⟩) (carrier_union _ _).le, ?_⟩ + simp_rw [carrier_union, iSup_range', eq] + exact (subset_adjoin _ _).trans (SetLike.coe_subset_coe.mpr <| + le_sup_right.trans <| le_iSup_of_le (Classical.arbitrary _) le_rfl) + +end Chain + +theorem nonempty_algHom_of_exist_lifts_finset [alg : Algebra.IsAlgebraic F E] + (h : ∀ S : Finset E, ∃ σ : Lifts F E K, (S : Set E) ⊆ σ.carrier) : + Nonempty (E →ₐ[F] K) := by + have : (⊥ : Lifts F E K).IsExtendible := fun S ↦ have ⟨σ, hσ⟩ := h S; ⟨σ, bot_le, hσ⟩ + have ⟨ϕ, hϕ⟩ := zorn_le₀ {ϕ : Lifts F E K | ϕ.IsExtendible} + fun c hext hc ↦ (isEmpty_or_nonempty c).elim + (fun _ ↦ ⟨⊥, this, fun ϕ hϕ ↦ isEmptyElim (⟨ϕ, hϕ⟩ : c)⟩) + fun _ ↦ ⟨_, union_isExtendible c hc hext, le_union c hc⟩ + suffices ϕ.carrier = ⊤ from ⟨ϕ.emb.comp <| ((equivOfEq this).trans topEquiv).symm⟩ + by_contra! + obtain ⟨α, -, hα⟩ := SetLike.exists_of_lt this.lt_top + let _ : Algebra ϕ.carrier K := ϕ.emb.toAlgebra + let Λ := ϕ.carrier⟮α⟯ →ₐ[ϕ.carrier] K + have := finiteDimensional_adjoin (S := {α}) fun _ _ ↦ ((alg.tower_top ϕ.carrier).isIntegral).1 _ + let L (σ : Λ) : Lifts F E K := ⟨ϕ.carrier⟮α⟯.restrictScalars F, σ.restrictScalars F⟩ + have hL (σ : Λ) : ϕ < L σ := lt_iff.mpr + ⟨by simpa only [restrictScalars_adjoin_eq_sup, left_lt_sup, adjoin_simple_le_iff], + AlgHom.coe_ringHom_injective σ.comp_algebraMap⟩ + have ⟨(ϕ_ext : ϕ.IsExtendible), ϕ_max⟩ := maximal_iff_forall_gt.mp hϕ + simp_rw [Set.mem_setOf, IsExtendible] at ϕ_max; push_neg at ϕ_max + choose S hS using fun σ : Λ ↦ ϕ_max (hL σ) + classical + have ⟨θ, hθϕ, hθ⟩ := ϕ_ext ({α} ∪ Finset.univ.biUnion S) + simp_rw [Finset.coe_union, Set.union_subset_iff, Finset.coe_singleton, Set.singleton_subset_iff, + Finset.coe_biUnion, Finset.coe_univ, Set.mem_univ, Set.iUnion_true, Set.iUnion_subset_iff] at hθ + have : ϕ.carrier⟮α⟯.restrictScalars F ≤ θ.carrier := by + rw [restrictScalars_adjoin_eq_sup, sup_le_iff, adjoin_simple_le_iff]; exact ⟨hθϕ.1, hθ.1⟩ + exact hS ⟨(θ.emb.comp <| inclusion this).toRingHom, hθϕ.2⟩ θ ⟨this, fun _ ↦ rfl⟩ (hθ.2 _) /-- Given a lift `x` and an integral element `s : E` over `x.carrier` whose conjugates over `x.carrier` are all in `K`, we can extend the lift to a lift whose carrier contains `s`. -/ -theorem Lifts.exists_lift_of_splits' (x : Lifts F E K) {s : E} (h1 : IsIntegral x.carrier s) +theorem exists_lift_of_splits' (x : Lifts F E K) {s : E} (h1 : IsIntegral x.carrier s) (h2 : (minpoly x.carrier s).Splits x.emb.toRingHom) : ∃ y, x ≤ y ∧ s ∈ y.carrier := have I2 := (minpoly.degree_pos h1).ne' letI : Algebra x.carrier K := x.emb.toRingHom.toAlgebra @@ -87,11 +215,13 @@ theorem Lifts.exists_lift_of_splits' (x : Lifts F E K) {s : E} (h1 : IsIntegral /-- Given an integral element `s : E` over `F` whose `F`-conjugates are all in `K`, any lift can be extended to one whose carrier contains `s`. -/ -theorem Lifts.exists_lift_of_splits (x : Lifts F E K) {s : E} (h1 : IsIntegral F s) +theorem exists_lift_of_splits (x : Lifts F E K) {s : E} (h1 : IsIntegral F s) (h2 : (minpoly F s).Splits (algebraMap F K)) : ∃ y, x ≤ y ∧ s ∈ y.carrier := - Lifts.exists_lift_of_splits' x h1.tower_top <| h1.minpoly_splits_tower_top' <| by + exists_lift_of_splits' x h1.tower_top <| h1.minpoly_splits_tower_top' <| by rwa [← x.emb.comp_algebraMap] at h2 +end Lifts + section private theorem exists_algHom_adjoin_of_splits'' {L : IntermediateField F E} diff --git a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean index c2bc06cd1f067..c164bdbac9521 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean @@ -32,10 +32,12 @@ polynomial in `k` splits. algebraic closure, algebraically closed -## TODO +## Main reults -- Prove that if `K / k` is algebraic, and any monic irreducible polynomial over `k` has a root - in `K`, then `K` is algebraically closed (in fact an algebraic closure of `k`). +- `IsAlgClosure.of_splits`: if `K / k` is algebraic, and every monic irreducible polynomial over + `k` splits in `K`, then `K` is algebraically closed (in fact an algebraic closure of `k`). + For the stronger fact that only requires every such polynomial has a root in `K`, + see `IsAlgClosure.of_exist_roots`. Reference: , Theorem 2 diff --git a/Mathlib/FieldTheory/Isaacs.lean b/Mathlib/FieldTheory/Isaacs.lean new file mode 100644 index 0000000000000..2219c66b3c027 --- /dev/null +++ b/Mathlib/FieldTheory/Isaacs.lean @@ -0,0 +1,115 @@ +/- +Copyright (c) 2024 Junyan Xu. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Junyan Xu +-/ +import Mathlib.FieldTheory.PrimitiveElement +import Mathlib.GroupTheory.CosetCover + +/-! +# Algebraic extensions are determined by their sets of minimal polynomials up to isomorphism + +## Main results + +`Field.nonempty_algHom_of_exist_roots` says if `E/F` and `K/F` are field extensions +with `E/F` algebraic, and if the minimal polynomial of every element of `E` over `F` has a root +in `K`, then there exists an `F`-embedding of `E` into `K`. If `E/F` and `K/F` have the same +set of minimal polynomials, then `E` and `K` are isomorphic as `F`-algebras. As a corollary: + +`IsAlgClosure.of_exist_roots`: if `E/F` is algebraic and every monic irreducible polynomial +in `F[X]` has a root in `E`, then `E` is an algebraic closure of `F`. + +## Reference + +[Isaacs1980] *Roots of Polynomials in Algebraic Extensions of Fields*, +The American Mathematical Monthly + +-/ + +namespace Field + +open Polynomial IntermediateField + +variable {F E K : Type*} [Field F] [Field E] [Field K] [Algebra F E] [Algebra F K] +variable [alg : Algebra.IsAlgebraic F E] + +theorem nonempty_algHom_of_exist_roots (h : ∀ x : E, ∃ y : K, aeval y (minpoly F x) = 0) : + Nonempty (E →ₐ[F] K) := by + refine Lifts.nonempty_algHom_of_exist_lifts_finset fun S ↦ ⟨⟨adjoin F S, ?_⟩, subset_adjoin _ _⟩ + let p := (S.prod <| minpoly F).map (algebraMap F K) + let K' := SplittingField p + have splits s (hs : s ∈ S) : (minpoly F s).Splits (algebraMap F K') := by + apply splits_of_splits_of_dvd _ + (Finset.prod_ne_zero_iff.mpr fun _ _ ↦ minpoly.ne_zero <| (alg.isIntegral).1 _) + ((splits_map_iff _ _).mp <| SplittingField.splits p) (Finset.dvd_prod_of_mem _ hs) + let K₀ := (⊥ : IntermediateField K K').restrictScalars F + let FS := adjoin F (S : Set E) + let Ω := FS →ₐ[F] K' + have := finiteDimensional_adjoin (S := (S : Set E)) fun _ _ ↦ (alg.isIntegral).1 _ + let M (ω : Ω) := Subalgebra.toSubmodule (K₀.comap ω).toSubalgebra + have : ⋃ ω : Ω, (M ω : Set FS) = Set.univ := + Set.eq_univ_of_forall fun ⟨α, hα⟩ ↦ Set.mem_iUnion.mpr <| by + have ⟨β, hβ⟩ := h α + let ϕ : F⟮α⟯ →ₐ[F] K' := (IsScalarTower.toAlgHom _ _ _).comp ((AdjoinRoot.liftHom _ _ hβ).comp + (adjoinRootEquivAdjoin F <| (alg.isIntegral).1 _).symm.toAlgHom) + have ⟨ω, hω⟩ := exists_algHom_adjoin_of_splits + (fun s hs ↦ ⟨(alg.isIntegral).1 _, splits s hs⟩) ϕ (adjoin_simple_le_iff.mpr hα) + refine ⟨ω, β, ((DFunLike.congr_fun hω <| AdjoinSimple.gen F α).trans ?_).symm⟩ + rw [AlgHom.comp_apply, AlgHom.comp_apply, AlgEquiv.coe_algHom, + adjoinRootEquivAdjoin_symm_apply_gen, AdjoinRoot.liftHom_root] + rfl + have ω : ∃ ω : Ω, ⊤ ≤ M ω := by + cases finite_or_infinite F + · have ⟨α, hα⟩ := exists_primitive_element_of_finite_bot F FS + have ⟨ω, hω⟩ := Set.mem_iUnion.mp (this ▸ Set.mem_univ α) + exact ⟨ω, show ⊤ ≤ K₀.comap ω by rwa [← hα, adjoin_simple_le_iff]⟩ + · simp_rw [top_le_iff, Subspace.exists_eq_top_of_iUnion_eq_univ this] + exact ((botEquiv K K').toAlgHom.restrictScalars F).comp + (ω.choose.codRestrict K₀.toSubalgebra fun x ↦ ω.choose_spec trivial) + +theorem nonempty_algHom_of_minpoly_eq + (h : ∀ x : E, ∃ y : K, minpoly F x = minpoly F y) : + Nonempty (E →ₐ[F] K) := + nonempty_algHom_of_exist_roots fun x ↦ have ⟨y, hy⟩ := h x; ⟨y, by rw [hy, minpoly.aeval]⟩ + +theorem nonempty_algHom_of_range_minpoly_subset + (h : Set.range (@minpoly F E _ _ _) ⊆ Set.range (@minpoly F K _ _ _)) : + Nonempty (E →ₐ[F] K) := + nonempty_algHom_of_minpoly_eq fun x ↦ have ⟨y, hy⟩ := h ⟨x, rfl⟩; ⟨y, hy.symm⟩ + +theorem nonempty_algEquiv_of_range_minpoly_eq + (h : Set.range (@minpoly F E _ _ _) = Set.range (@minpoly F K _ _ _)) : + Nonempty (E ≃ₐ[F] K) := + have ⟨σ⟩ := nonempty_algHom_of_range_minpoly_subset h.le + have : Algebra.IsAlgebraic F K := ⟨fun y ↦ IsIntegral.isAlgebraic <| by + by_contra hy + have ⟨x, hx⟩ := h.ge ⟨y, rfl⟩ + rw [minpoly.eq_zero hy] at hx + exact minpoly.ne_zero ((alg.isIntegral).1 x) hx⟩ + have ⟨τ⟩ := nonempty_algHom_of_range_minpoly_subset h.ge + ⟨.ofBijective _ (Algebra.IsAlgebraic.algHom_bijective₂ σ τ).1⟩ + +theorem nonempty_algHom_of_aeval_eq_zero_subset + (h : {p : F[X] | ∃ x : E, aeval x p = 0} ⊆ {p | ∃ y : K, aeval y p = 0}) : + Nonempty (E →ₐ[F] K) := + nonempty_algHom_of_minpoly_eq fun x ↦ + have ⟨y, hy⟩ := h ⟨_, minpoly.aeval F x⟩ + ⟨y, (minpoly.eq_iff_aeval_minpoly_eq_zero <| (alg.isIntegral).1 x).mpr hy⟩ + +theorem nonempty_algEquiv_of_aeval_eq_zero_eq [Algebra.IsAlgebraic F K] + (h : {p : F[X] | ∃ x : E, aeval x p = 0} = {p | ∃ y : K, aeval y p = 0}) : + Nonempty (E ≃ₐ[F] K) := + have ⟨σ⟩ := nonempty_algHom_of_aeval_eq_zero_subset h.le + have ⟨τ⟩ := nonempty_algHom_of_aeval_eq_zero_subset h.ge + ⟨.ofBijective _ (Algebra.IsAlgebraic.algHom_bijective₂ σ τ).1⟩ + +theorem _root_.IsAlgClosure.of_exist_roots + (h : ∀ p : F[X], p.Monic → Irreducible p → ∃ x : E, aeval x p = 0) : + IsAlgClosure F E := + .of_splits fun p _ _ ↦ + have ⟨σ⟩ := nonempty_algHom_of_exist_roots fun x : p.SplittingField ↦ + have := Algebra.IsAlgebraic.isIntegral (K := F).1 x + h _ (minpoly.monic this) (minpoly.irreducible this) + splits_of_algHom (SplittingField.splits _) σ + +end Field diff --git a/docs/references.bib b/docs/references.bib index 392120a19de1e..f227b4270fd17 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -2083,6 +2083,19 @@ @Book{ IrelandRosen1990 url = {https://doi.org/10.1007/978-1-4757-2103-4} } +@Article{ Isaacs1980, + author = {Isaacs, I. M.}, + title = {Roots of Polynomials in Algebraic Extensions of Fields}, + publisher = {Taylor \& Francis}, + year = {1980}, + pages = {543--544}, + journal = {The American Mathematical Monthly}, + volume = {87}, + number = {7}, + doi = {10.1080/00029890.1980.11995085}, + url = {https://doi.org/10.1080/00029890.1980.11995085} +} + @Book{ iversen, title = {Generic Local Structure of the Morphisms in Commutative Algebra}, From 1ef6e445ac54d2bb1d2568e3bc000b8512a0a765 Mon Sep 17 00:00:00 2001 From: Hannah Scholz Date: Sat, 23 Nov 2024 06:30:17 +0000 Subject: [PATCH 220/829] feat: image of any interval under `affineHomeomorph` (#19301) --- Mathlib/Data/Set/Pointwise/Interval.lean | 37 ++++++++++++++++++++++++ Mathlib/Topology/Algebra/Field.lean | 21 ++++++++++++++ 2 files changed, 58 insertions(+) diff --git a/Mathlib/Data/Set/Pointwise/Interval.lean b/Mathlib/Data/Set/Pointwise/Interval.lean index 535b6e9100c34..e6846ddbc75b3 100644 --- a/Mathlib/Data/Set/Pointwise/Interval.lean +++ b/Mathlib/Data/Set/Pointwise/Interval.lean @@ -725,6 +725,22 @@ theorem image_mul_left_Ioo {a : α} (h : 0 < a) (b c : α) : (a * ·) '' Ioo b c = Ioo (a * b) (a * c) := by convert image_mul_right_Ioo b c h using 1 <;> simp only [mul_comm _ a] +theorem image_mul_right_Ico (a b : α) {c : α} (h : 0 < c) : + (fun x => x * c) '' Ico a b = Ico (a * c) (b * c) := + ((Units.mk0 c h.ne').mulRight.image_eq_preimage _).trans (by simp [h, division_def]) + +theorem image_mul_left_Ico {a : α} (h : 0 < a) (b c : α) : + (a * ·) '' Ico b c = Ico (a * b) (a * c) := by + convert image_mul_right_Ico b c h using 1 <;> simp only [mul_comm _ a] + +theorem image_mul_right_Ioc (a b : α) {c : α} (h : 0 < c) : + (fun x => x * c) '' Ioc a b = Ioc (a * c) (b * c) := + ((Units.mk0 c h.ne').mulRight.image_eq_preimage _).trans (by simp [h, division_def]) + +theorem image_mul_left_Ioc {a : α} (h : 0 < a) (b c : α) : + (a * ·) '' Ioc b c = Ioc (a * b) (a * c) := by + convert image_mul_right_Ioc b c h using 1 <;> simp only [mul_comm _ a] + /-- The (pre)image under `inv` of `Ioo 0 a` is `Ioi a⁻¹`. -/ theorem inv_Ioo_0_left {a : α} (ha : 0 < a) : (Ioo 0 a)⁻¹ = Ioi a⁻¹ := by ext x @@ -754,6 +770,27 @@ theorem image_affine_Icc' {a : α} (h : 0 < a) (b c d : α) : rwa [Set.image_image] at this rw [image_mul_left_Icc' h, image_add_const_Icc] +@[simp] +theorem image_affine_Ico {a : α} (h : 0 < a) (b c d : α) : + (a * · + b) '' Ico c d = Ico (a * c + b) (a * d + b) := by + suffices (· + b) '' ((a * ·) '' Ico c d) = Ico (a * c + b) (a * d + b) by + rwa [Set.image_image] at this + rw [image_mul_left_Ico h, image_add_const_Ico] + +@[simp] +theorem image_affine_Ioc {a : α} (h : 0 < a) (b c d : α) : + (a * · + b) '' Ioc c d = Ioc (a * c + b) (a * d + b) := by + suffices (· + b) '' ((a * ·) '' Ioc c d) = Ioc (a * c + b) (a * d + b) by + rwa [Set.image_image] at this + rw [image_mul_left_Ioc h, image_add_const_Ioc] + +@[simp] +theorem image_affine_Ioo {a : α} (h : 0 < a) (b c d : α) : + (a * · + b) '' Ioo c d = Ioo (a * c + b) (a * d + b) := by + suffices (· + b) '' ((a * ·) '' Ioo c d) = Ioo (a * c + b) (a * d + b) by + rwa [Set.image_image] at this + rw [image_mul_left_Ioo h, image_add_const_Ioo] + end LinearOrderedField end Set diff --git a/Mathlib/Topology/Algebra/Field.lean b/Mathlib/Topology/Algebra/Field.lean index 48b1183bc9f01..7e78ef983d7f6 100644 --- a/Mathlib/Topology/Algebra/Field.lean +++ b/Mathlib/Topology/Algebra/Field.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.GroupWithZero.Divisibility import Mathlib.Topology.Algebra.GroupWithZero import Mathlib.Topology.Algebra.Ring.Basic import Mathlib.Topology.Order.LocalExtr +import Mathlib.Data.Set.Pointwise.Interval /-! # Topological fields @@ -91,6 +92,26 @@ def affineHomeomorph (a b : 𝕜) (h : a ≠ 0) : 𝕜 ≃ₜ 𝕜 where exact mul_div_cancel_left₀ x h right_inv y := by simp [mul_div_cancel₀ _ h] +theorem affineHomeomorph_image_Icc {𝕜 : Type*} [LinearOrderedField 𝕜] [TopologicalSpace 𝕜] + [TopologicalRing 𝕜] (a b c d : 𝕜) (h : 0 < a) : + affineHomeomorph a b h.ne' '' Set.Icc c d = Set.Icc (a * c + b) (a * d + b) := by + simp [h] + +theorem affineHomeomorph_image_Ico {𝕜 : Type*} [LinearOrderedField 𝕜] [TopologicalSpace 𝕜] + [TopologicalRing 𝕜] (a b c d : 𝕜) (h : 0 < a) : + affineHomeomorph a b h.ne' '' Set.Ico c d = Set.Ico (a * c + b) (a * d + b) := by + simp [h] + +theorem affineHomeomorph_image_Ioc {𝕜 : Type*} [LinearOrderedField 𝕜] [TopologicalSpace 𝕜] + [TopologicalRing 𝕜] (a b c d : 𝕜) (h : 0 < a) : + affineHomeomorph a b h.ne' '' Set.Ioc c d = Set.Ioc (a * c + b) (a * d + b) := by + simp [h] + +theorem affineHomeomorph_image_Ioo {𝕜 : Type*} [LinearOrderedField 𝕜] [TopologicalSpace 𝕜] + [TopologicalRing 𝕜] (a b c d : 𝕜) (h : 0 < a) : + affineHomeomorph a b h.ne' '' Set.Ioo c d = Set.Ioo (a * c + b) (a * d + b) := by + simp [h] + end affineHomeomorph section LocalExtr From 02990e46811c8ead901b6c056cf91dc7e3e80c10 Mon Sep 17 00:00:00 2001 From: blizzard_inc Date: Sat, 23 Nov 2024 06:30:18 +0000 Subject: [PATCH 221/829] chore(Algebra/Module/Torsion): change adaptation note links to lean4 issue #1910 to refer to lean4 issue #1629 instead (#19381) These notes refer to the fact that `LinearMap.ker` and friends cannot be used as `f.ker`. However, this fails not because the function `LinearMap.ker` needs to be coerced to a proper function (which the issue leanprover/lean4#1910 is about), but because the argument `f` has a type (`LinearMap`) that doesn't occur directly as type of an argument to `LinearMap.ker`. The issue leanprover/lean4#1629 *does* suggest a fix for this. --- Mathlib/Algebra/Module/Torsion.lean | 4 ++-- Mathlib/LinearAlgebra/Dual.lean | 36 ++++++++++++++--------------- 2 files changed, 20 insertions(+), 20 deletions(-) diff --git a/Mathlib/Algebra/Module/Torsion.lean b/Mathlib/Algebra/Module/Torsion.lean index a993b226e188a..3121a844257b9 100644 --- a/Mathlib/Algebra/Module/Torsion.lean +++ b/Mathlib/Algebra/Module/Torsion.lean @@ -67,7 +67,7 @@ variable (R M : Type*) [Semiring R] [AddCommMonoid M] [Module R M] /-- The torsion ideal of `x`, containing all `a` such that `a • x = 0`. -/ @[simps!] def torsionOf (x : M) : Ideal R := - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation on LinearMap.ker https://github.com/leanprover/lean4/issues/1910 + -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation on LinearMap.ker https://github.com/leanprover/lean4/issues/1629 LinearMap.ker (LinearMap.toSpanSingleton R M x) @[simp] @@ -148,7 +148,7 @@ namespace Submodule `a • x = 0`. -/ @[simps!] def torsionBy (a : R) : Submodule R M := - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation on LinearMap.ker https://github.com/leanprover/lean4/issues/1910 + -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation on LinearMap.ker https://github.com/leanprover/lean4/issues/1629 LinearMap.ker (DistribMulAction.toLinearMap R M a) /-- The submodule containing all elements `x` of `M` such that `a • x = 0` for all `a` in `s`. -/ diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index 1ef8aede9c87b..a203905b0a8a2 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -339,11 +339,11 @@ theorem toDual_injective : Injective b.toDual := fun x y h ↦ b.ext_elem_iff.mp theorem toDual_inj (m : M) (a : b.toDual m = 0) : m = 0 := b.toDual_injective (by rwa [_root_.map_zero]) --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.ker +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.ker theorem toDual_ker : LinearMap.ker b.toDual = ⊥ := ker_eq_bot'.mpr b.toDual_inj --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.range +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.range theorem toDual_range [Finite ι] : LinearMap.range b.toDual = ⊤ := by refine eq_top_iff'.2 fun f => ?_ let lin_comb : ι →₀ R := Finsupp.equivFunOnFinite.symm fun i => f (b i) @@ -447,7 +447,7 @@ theorem eval_ker {ι : Type*} (b : Basis ι R M) : simp_rw [LinearMap.ext_iff, Dual.eval_apply, zero_apply] at hm exact (Basis.forall_coord_eq_zero_iff _).mp fun i => hm (b.coord i) --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.range +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.range theorem eval_range {ι : Type*} [Finite ι] (b : Basis ι R M) : LinearMap.range (Dual.eval R M) = ⊤ := by classical @@ -501,7 +501,7 @@ section variable (K) (V) --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.ker +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.ker theorem eval_ker : LinearMap.ker (eval K V) = ⊥ := have ⟨s, hs⟩ := Module.projective_def'.mp ‹Projective K V› ker_eq_bot.mpr <| .of_comp (f := s.dualMap.dualMap) <| (ker_eq_bot.mp <| @@ -595,7 +595,7 @@ instance (priority := 900) IsReflexive.of_finite_of_free [Module.Finite R M] [Fr variable [IsReflexive R M] --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.range +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.range theorem erange_coe : LinearMap.range (eval R M) = ⊤ := range_eq_top.mpr (bijective_dual_eval _ _).2 @@ -943,7 +943,7 @@ theorem dualRestrict_apply (W : Submodule R M) (φ : Module.Dual R M) (x : W) : that `φ w = 0` for all `w ∈ W`. -/ def dualAnnihilator {R : Type u} {M : Type v} [CommSemiring R] [AddCommMonoid M] [Module R M] (W : Submodule R M) : Submodule R <| Module.Dual R M := --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.ker +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.ker LinearMap.ker W.dualRestrict @[simp] @@ -955,7 +955,7 @@ theorem mem_dualAnnihilator (φ : Module.Dual R M) : φ ∈ W.dualAnnihilator /-- That $\operatorname{ker}(\iota^* : V^* \to W^*) = \operatorname{ann}(W)$. This is the definition of the dual annihilator of the submodule $W$. -/ theorem dualRestrict_ker_eq_dualAnnihilator (W : Submodule R M) : - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.ker + -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.ker LinearMap.ker W.dualRestrict = W.dualAnnihilator := rfl @@ -1189,7 +1189,7 @@ theorem quotAnnihilatorEquiv_apply (W : Subspace K V) (φ : Module.Dual K V) : rfl /-- The natural isomorphism from the dual of a subspace `W` to `W.dualLift.range`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.range +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.range noncomputable def dualEquivDual (W : Subspace K V) : Module.Dual K W ≃ₗ[K] LinearMap.range W.dualLift := LinearEquiv.ofInjective _ dualLift_injective @@ -1228,7 +1228,7 @@ theorem dualAnnihilator_dualAnnihilator_eq (W : Subspace K V) : rwa [← OrderIso.symm_apply_eq] /-- The quotient by the dual is isomorphic to its dual annihilator. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.range +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.range noncomputable def quotDualEquivAnnihilator (W : Subspace K V) : (Module.Dual K V ⧸ LinearMap.range W.dualLift) ≃ₗ[K] W.dualAnnihilator := LinearEquiv.quotEquivOfQuotEquiv <| LinearEquiv.trans W.quotAnnihilatorEquiv W.dualEquivDual @@ -1274,7 +1274,7 @@ variable {R : Type uR} [CommSemiring R] {M₁ : Type uM₁} {M₂ : Type uM₂} variable [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] variable (f : M₁ →ₗ[R] M₂) --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.ker +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.ker theorem ker_dualMap_eq_dualAnnihilator_range : LinearMap.ker f.dualMap = f.range.dualAnnihilator := by ext @@ -1282,7 +1282,7 @@ theorem ker_dualMap_eq_dualAnnihilator_range : ← SetLike.mem_coe, range_coe, Set.forall_mem_range] rfl --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.range +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.range theorem range_dualMap_le_dualAnnihilator_ker : LinearMap.range f.dualMap ≤ f.ker.dualAnnihilator := by rintro _ ⟨ψ, rfl⟩ @@ -1338,7 +1338,7 @@ theorem dualPairing_apply {W : Submodule R M} (φ : Module.Dual R M) (x : W) : W.dualPairing (Quotient.mk φ) x = φ x := rfl --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.range +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.range /-- That $\operatorname{im}(q^* : (V/W)^* \to V^*) = \operatorname{ann}(W)$. -/ theorem range_dualMap_mkQ_eq (W : Submodule R M) : LinearMap.range W.mkQ.dualMap = W.dualAnnihilator := by @@ -1360,7 +1360,7 @@ def dualQuotEquivDualAnnihilator (W : Submodule R M) : Module.Dual R (M ⧸ W) ≃ₗ[R] W.dualAnnihilator := LinearEquiv.ofLinear (W.mkQ.dualMap.codRestrict W.dualAnnihilator fun φ => --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.mem_range_self +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.mem_range_self W.range_dualMap_mkQ_eq ▸ LinearMap.mem_range_self W.mkQ.dualMap φ) W.dualCopairing (by ext; rfl) (by ext; rfl) @@ -1415,7 +1415,7 @@ namespace LinearMap open Submodule --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.range +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.range theorem range_dualMap_eq_dualAnnihilator_ker_of_surjective (f : M →ₗ[R] M') (hf : Function.Surjective f) : LinearMap.range f.dualMap = f.ker.dualAnnihilator := ((f.quotKerEquivOfSurjective hf).dualMap.range_comp _).trans f.ker.range_dualMap_mkQ_eq @@ -1429,7 +1429,7 @@ theorem range_dualMap_eq_dualAnnihilator_ker_of_subtype_range_surjective (f : M rw [← range_eq_top, range_rangeRestrict] have := range_dualMap_eq_dualAnnihilator_ker_of_surjective f.rangeRestrict rr_surj convert this using 1 - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 + -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 · calc _ = range ((range f).subtype.comp f.rangeRestrict).dualMap := by simp _ = _ := ?_ @@ -1525,7 +1525,7 @@ theorem dualMap_surjective_of_injective {f : V₁ →ₗ[K] V₂} (hf : Function have ⟨f', hf'⟩ := f.exists_leftInverse_of_injective (ker_eq_bot.mpr hf) ⟨φ.comp f', ext fun x ↦ congr(φ <| $hf' x)⟩ - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.range + -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.range theorem range_dualMap_eq_dualAnnihilator_ker (f : V₁ →ₗ[K] V₂) : LinearMap.range f.dualMap = f.ker.dualAnnihilator := range_dualMap_eq_dualAnnihilator_ker_of_subtype_range_surjective f <| @@ -1576,7 +1576,7 @@ theorem dualAnnihilator_inf_eq (W W' : Subspace K V₁) : (W ⊓ W').dualAnnihilator = W.dualAnnihilator ⊔ W'.dualAnnihilator := by refine le_antisymm ?_ (sup_dualAnnihilator_le_inf W W') let F : V₁ →ₗ[K] (V₁ ⧸ W) × V₁ ⧸ W' := (Submodule.mkQ W).prod (Submodule.mkQ W') - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.ker + -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.ker have : LinearMap.ker F = W ⊓ W' := by simp only [F, LinearMap.ker_prod, ker_mkQ] rw [← this, ← LinearMap.range_dualMap_eq_dualAnnihilator_ker] intro φ @@ -1632,7 +1632,7 @@ namespace LinearMap @[simp] theorem finrank_range_dualMap_eq_finrank_range (f : V₁ →ₗ[K] V₂) : - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 + -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 finrank K (LinearMap.range f.dualMap) = finrank K (LinearMap.range f) := by rw [congr_arg dualMap (show f = (range f).subtype.comp f.rangeRestrict by rfl), ← dualMap_comp_dualMap, range_comp, From ff064a56e58e8ceb1c0a3febd60e701a1cf408fe Mon Sep 17 00:00:00 2001 From: blizzard_inc Date: Sat, 23 Nov 2024 06:30:20 +0000 Subject: [PATCH 222/829] chore: change `pp_nodot` adaptation note links to lean4 issue #1910 to refer to lean4 issue #6178 instead (#19385) These adaptation notes refer to the `pp_nodot` attribute not doing anything on certain declarations, due to their inability to be used with dot notation. However, now that leanprover/lean4#1910 has been fixed, the limiting factor has become the fact that the pretty printer doesn't work for these kinds of dot notation, which leanprover/lean4#6178 is an issue for. As such, the links to the blocking lean issue have been changed in this PR. --- Mathlib/Algebra/Symmetrized.lean | 2 +- Mathlib/Data/NNReal/Defs.lean | 2 +- Mathlib/Data/Real/Sqrt.lean | 2 +- 3 files changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/Algebra/Symmetrized.lean b/Mathlib/Algebra/Symmetrized.lean index 00c398aaf91dd..45a69d0768566 100644 --- a/Mathlib/Algebra/Symmetrized.lean +++ b/Mathlib/Algebra/Symmetrized.lean @@ -48,7 +48,7 @@ def sym : α ≃ αˢʸᵐ := /-- The element of `α` represented by `x : αˢʸᵐ`. -/ -- Porting note (kmill): `pp_nodot` has no affect here --- unless RFC https://github.com/leanprover/lean4/issues/1910 leads to dot notation for CoeFun +-- unless RFC https://github.com/leanprover/lean4/issues/6178 leads to dot notation pp for CoeFun @[pp_nodot] def unsym : αˢʸᵐ ≃ α := Equiv.refl _ diff --git a/Mathlib/Data/NNReal/Defs.lean b/Mathlib/Data/NNReal/Defs.lean index 210cfbe55782c..9c117c2ed6f81 100644 --- a/Mathlib/Data/NNReal/Defs.lean +++ b/Mathlib/Data/NNReal/Defs.lean @@ -917,7 +917,7 @@ namespace Real /-- The absolute value on `ℝ` as a map to `ℝ≥0`. -/ -- Porting note (kmill): `pp_nodot` has no affect here --- unless RFC https://github.com/leanprover/lean4/issues/1910 leads to dot notation for CoeFun +-- unless RFC https://github.com/leanprover/lean4/issues/6178 leads to dot notation pp for CoeFun @[pp_nodot] def nnabs : ℝ →*₀ ℝ≥0 where toFun x := ⟨|x|, abs_nonneg x⟩ diff --git a/Mathlib/Data/Real/Sqrt.lean b/Mathlib/Data/Real/Sqrt.lean index b07597c19deda..b43bd46488f69 100644 --- a/Mathlib/Data/Real/Sqrt.lean +++ b/Mathlib/Data/Real/Sqrt.lean @@ -39,7 +39,7 @@ variable {x y : ℝ≥0} /-- Square root of a nonnegative real number. -/ -- Porting note (kmill): `pp_nodot` has no affect here --- unless RFC https://github.com/leanprover/lean4/issues/1910 leads to dot notation for CoeFun +-- unless RFC https://github.com/leanprover/lean4/issues/6178 leads to dot notation pp for CoeFun @[pp_nodot] noncomputable def sqrt : ℝ≥0 ≃o ℝ≥0 := OrderIso.symm <| powOrderIso 2 two_ne_zero From 4b6c6f8752e065dc0040b8165cfb7f76f1e87467 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 23 Nov 2024 06:55:58 +0000 Subject: [PATCH 223/829] chore: cleanup of nolint simpNF (#19061) --- Mathlib/Algebra/Algebra/Unitization.lean | 3 +-- Mathlib/Algebra/BigOperators/Finsupp.lean | 14 +++++--------- Mathlib/Algebra/Group/WithOne/Basic.lean | 3 +-- .../Algebra/Homology/ShortComplex/Homology.lean | 2 -- .../Homology/ShortComplex/LeftHomology.lean | 1 - .../Homology/ShortComplex/RightHomology.lean | 1 - Mathlib/Algebra/Lie/IdealOperations.lean | 2 -- Mathlib/Algebra/Module/GradedModule.lean | 4 +--- Mathlib/Algebra/Module/Submodule/Pointwise.lean | 4 +--- Mathlib/Algebra/MvPolynomial/Basic.lean | 5 ++--- Mathlib/Algebra/Ring/Semiconj.lean | 8 ++------ Mathlib/Algebra/Star/Subsemiring.lean | 1 - Mathlib/RingTheory/DedekindDomain/Ideal.lean | 4 ---- 13 files changed, 13 insertions(+), 39 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Unitization.lean b/Mathlib/Algebra/Algebra/Unitization.lean index 42a25ab336b19..89f706b546d03 100644 --- a/Mathlib/Algebra/Algebra/Unitization.lean +++ b/Mathlib/Algebra/Algebra/Unitization.lean @@ -743,8 +743,7 @@ def starLift : (A →⋆ₙₐ[R] C) ≃ (Unitization R A →⋆ₐ[R] C) := right_inv := fun φ => Unitization.algHom_ext'' <| by simp } --- Note (#6057) : tagging simpNF because linter complains -@[simp high, nolint simpNF] +@[simp high] theorem starLift_symm_apply_apply (φ : Unitization R A →⋆ₐ[R] C) (a : A) : Unitization.starLift.symm φ a = φ a := rfl diff --git a/Mathlib/Algebra/BigOperators/Finsupp.lean b/Mathlib/Algebra/BigOperators/Finsupp.lean index a02b2d8a48142..dffd8f94b48be 100644 --- a/Mathlib/Algebra/BigOperators/Finsupp.lean +++ b/Mathlib/Algebra/BigOperators/Finsupp.lean @@ -86,22 +86,17 @@ theorem prod_ite_eq [DecidableEq α] (f : α →₀ M) (a : α) (b : α → M dsimp [Finsupp.prod] rw [f.support.prod_ite_eq] -/- Porting note: simpnf linter, added aux lemma below -Left-hand side simplifies from - Finsupp.sum f fun x v => if a = x then v else 0 -to - if ↑f a = 0 then 0 else ↑f a --/ --- @[simp] theorem sum_ite_self_eq [DecidableEq α] {N : Type*} [AddCommMonoid N] (f : α →₀ N) (a : α) : (f.sum fun x v => ite (a = x) v 0) = f a := by classical convert f.sum_ite_eq a fun _ => id simp [ite_eq_right_iff.2 Eq.symm] --- Porting note: Added this thm to replace the simp in the previous one. Need to add [DecidableEq N] +/-- +The left hand side of `sum_ite_self_eq` simplifies; this is the variant that is useful for `simp`. +-/ @[simp] -theorem sum_ite_self_eq_aux [DecidableEq α] {N : Type*} [AddCommMonoid N] (f : α →₀ N) (a : α) : +theorem if_mem_support [DecidableEq α] {N : Type*} [AddCommMonoid N] (f : α →₀ N) (a : α) : (if a ∈ f.support then f a else 0) = f a := by simp only [mem_support_iff, ne_eq, ite_eq_left_iff, not_not] exact fun h ↦ h.symm @@ -113,6 +108,7 @@ theorem prod_ite_eq' [DecidableEq α] (f : α →₀ M) (a : α) (b : α → M dsimp [Finsupp.prod] rw [f.support.prod_ite_eq'] +/-- A restatement of `sum_ite_self_eq` with the equality test reversed. -/ theorem sum_ite_self_eq' [DecidableEq α] {N : Type*} [AddCommMonoid N] (f : α →₀ N) (a : α) : (f.sum fun x v => ite (x = a) v 0) = f a := by classical diff --git a/Mathlib/Algebra/Group/WithOne/Basic.lean b/Mathlib/Algebra/Group/WithOne/Basic.lean index 3dca5ba0dab77..b9d87ce1b04af 100644 --- a/Mathlib/Algebra/Group/WithOne/Basic.lean +++ b/Mathlib/Algebra/Group/WithOne/Basic.lean @@ -73,8 +73,7 @@ variable (f : α →ₙ* β) theorem lift_coe (x : α) : lift f x = f x := rfl --- Porting note (#11119): removed `simp` attribute to appease `simpNF` linter. -@[to_additive] +@[to_additive (attr := simp)] theorem lift_one : lift f 1 = 1 := rfl diff --git a/Mathlib/Algebra/Homology/ShortComplex/Homology.lean b/Mathlib/Algebra/Homology/ShortComplex/Homology.lean index 25854245e340c..0c416a49f4df5 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/Homology.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/Homology.lean @@ -70,8 +70,6 @@ structure HomologyMapData where namespace HomologyMapData -attribute [nolint simpNF] mk.injEq - variable {φ h₁ h₂} @[reassoc] diff --git a/Mathlib/Algebra/Homology/ShortComplex/LeftHomology.lean b/Mathlib/Algebra/Homology/ShortComplex/LeftHomology.lean index 1022a52c7ab90..a8abcedd4972c 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/LeftHomology.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/LeftHomology.lean @@ -266,7 +266,6 @@ structure LeftHomologyMapData where namespace LeftHomologyMapData attribute [reassoc (attr := simp)] commi commf' commπ -attribute [nolint simpNF] mk.injEq /-- The left homology map data associated to the zero morphism between two short complexes. -/ @[simps] diff --git a/Mathlib/Algebra/Homology/ShortComplex/RightHomology.lean b/Mathlib/Algebra/Homology/ShortComplex/RightHomology.lean index 865ffa00b94fb..3592289533893 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/RightHomology.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/RightHomology.lean @@ -356,7 +356,6 @@ structure RightHomologyMapData where namespace RightHomologyMapData attribute [reassoc (attr := simp)] commp commg' commι -attribute [nolint simpNF] mk.injEq /-- The right homology map data associated to the zero morphism between two short complexes. -/ @[simps] diff --git a/Mathlib/Algebra/Lie/IdealOperations.lean b/Mathlib/Algebra/Lie/IdealOperations.lean index 75e9facacd878..459dce3dfdbf0 100644 --- a/Mathlib/Algebra/Lie/IdealOperations.lean +++ b/Mathlib/Algebra/Lie/IdealOperations.lean @@ -192,12 +192,10 @@ theorem sup_lie : ⁅I ⊔ J, N⁆ = ⁅I, N⁆ ⊔ ⁅J, N⁆ := by use ⁅((⟨x₂, hx₂⟩ : J) : L), (n : N)⁆; constructor; · apply lie_coe_mem_lie simp [← h, ← hx'] --- @[simp] -- Porting note: not in simpNF theorem lie_inf : ⁅I, N ⊓ N'⁆ ≤ ⁅I, N⁆ ⊓ ⁅I, N'⁆ := by rw [le_inf_iff]; constructor <;> apply mono_lie_right <;> [exact inf_le_left; exact inf_le_right] --- @[simp] -- Porting note: not in simpNF theorem inf_lie : ⁅I ⊓ J, N⁆ ≤ ⁅I, N⁆ ⊓ ⁅J, N⁆ := by rw [le_inf_iff]; constructor <;> apply mono_lie_left <;> [exact inf_le_left; exact inf_le_right] diff --git a/Mathlib/Algebra/Module/GradedModule.lean b/Mathlib/Algebra/Module/GradedModule.lean index 8adb7c6ada0a5..8797f2bbdd7b5 100644 --- a/Mathlib/Algebra/Module/GradedModule.lean +++ b/Mathlib/Algebra/Module/GradedModule.lean @@ -93,11 +93,9 @@ theorem smulAddMonoidHom_apply_of_of [DecidableEq ιA] [DecidableEq ιB] [GMonoi smulAddMonoidHom A M (DirectSum.of A i x) (of M j y) = of M (i +ᵥ j) (GSMul.smul x y) := by simp [smulAddMonoidHom] --- @[simp] -- Porting note: simpNF lint theorem of_smul_of [DecidableEq ιA] [DecidableEq ιB] [GMonoid A] [Gmodule A M] {i j} (x : A i) (y : M j) : - DirectSum.of A i x • of M j y = of M (i +ᵥ j) (GSMul.smul x y) := - smulAddMonoidHom_apply_of_of _ _ _ _ + DirectSum.of A i x • of M j y = of M (i +ᵥ j) (GSMul.smul x y) := by simp open AddMonoidHom diff --git a/Mathlib/Algebra/Module/Submodule/Pointwise.lean b/Mathlib/Algebra/Module/Submodule/Pointwise.lean index ec25efdc44f71..99f269138c452 100644 --- a/Mathlib/Algebra/Module/Submodule/Pointwise.lean +++ b/Mathlib/Algebra/Module/Submodule/Pointwise.lean @@ -161,9 +161,7 @@ instance pointwiseAddCommMonoid : AddCommMonoid (Submodule R M) where theorem add_eq_sup (p q : Submodule R M) : p + q = p ⊔ q := rfl --- dsimp loops when applying this lemma to its LHS, --- probably https://github.com/leanprover/lean4/pull/2867 -@[simp, nolint simpNF] +@[simp] theorem zero_eq_bot : (0 : Submodule R M) = ⊥ := rfl diff --git a/Mathlib/Algebra/MvPolynomial/Basic.lean b/Mathlib/Algebra/MvPolynomial/Basic.lean index dac73ae71fe7a..b73915560d966 100644 --- a/Mathlib/Algebra/MvPolynomial/Basic.lean +++ b/Mathlib/Algebra/MvPolynomial/Basic.lean @@ -844,9 +844,8 @@ theorem constantCoeff_X (i : σ) : constantCoeff (X i : MvPolynomial σ R) = 0 : simp [constantCoeff_eq] variable {R} -/- porting note: increased priority because otherwise `simp` time outs when trying to simplify -the left-hand side. `simpNF` linter indicated this and it was verified. -/ -@[simp 1001] + +@[simp] theorem constantCoeff_smul {R : Type*} [SMulZeroClass R S₁] (a : R) (f : MvPolynomial σ S₁) : constantCoeff (a • f) = a • constantCoeff f := rfl diff --git a/Mathlib/Algebra/Ring/Semiconj.lean b/Mathlib/Algebra/Ring/Semiconj.lean index 8b20e68a71a06..0ec9636ddebc6 100644 --- a/Mathlib/Algebra/Ring/Semiconj.lean +++ b/Mathlib/Algebra/Ring/Semiconj.lean @@ -61,13 +61,9 @@ section variable [MulOneClass R] [HasDistribNeg R] --- Porting note: `simpNF` told me to remove `simp` attribute -theorem neg_one_right (a : R) : SemiconjBy a (-1) (-1) := - (one_right a).neg_right +theorem neg_one_right (a : R) : SemiconjBy a (-1) (-1) := by simp --- Porting note: `simpNF` told me to remove `simp` attribute -theorem neg_one_left (x : R) : SemiconjBy (-1) x x := - (SemiconjBy.one_left x).neg_left +theorem neg_one_left (x : R) : SemiconjBy (-1) x x := by simp end diff --git a/Mathlib/Algebra/Star/Subsemiring.lean b/Mathlib/Algebra/Star/Subsemiring.lean index 8ef2bdd8fdd39..707ac11a068a8 100644 --- a/Mathlib/Algebra/Star/Subsemiring.lean +++ b/Mathlib/Algebra/Star/Subsemiring.lean @@ -53,7 +53,6 @@ instance starRing (s : StarSubsemiring R) : StarRing s := instance semiring (s : StarSubsemiring R) : NonAssocSemiring s := s.toSubsemiring.toNonAssocSemiring -@[simp, nolint simpNF] theorem mem_carrier {s : StarSubsemiring R} {x : R} : x ∈ s.carrier ↔ x ∈ s := Iff.rfl diff --git a/Mathlib/RingTheory/DedekindDomain/Ideal.lean b/Mathlib/RingTheory/DedekindDomain/Ideal.lean index 22365ec289413..2533459ec4771 100644 --- a/Mathlib/RingTheory/DedekindDomain/Ideal.lean +++ b/Mathlib/RingTheory/DedekindDomain/Ideal.lean @@ -129,7 +129,6 @@ open Submodule Submodule.IsPrincipal theorem spanSingleton_inv (x : K) : (spanSingleton R₁⁰ x)⁻¹ = spanSingleton _ x⁻¹ := one_div_spanSingleton x --- @[simp] -- Porting note: not in simpNF form theorem spanSingleton_div_spanSingleton (x y : K) : spanSingleton R₁⁰ x / spanSingleton R₁⁰ y = spanSingleton R₁⁰ (x / y) := by rw [div_spanSingleton, mul_comm, spanSingleton_mul_spanSingleton, div_eq_mul_inv] @@ -1058,7 +1057,6 @@ variable [IsDedekindDomain R] (f : R ⧸ I ≃+* A ⧸ J) /-- The bijection between ideals of `R` dividing `I` and the ideals of `A` dividing `J` induced by an isomorphism `f : R/I ≅ A/J`. -/ --- @[simps] -- Porting note: simpNF complains about the lemmas generated by simps def idealFactorsEquivOfQuotEquiv : { p : Ideal R | p ∣ I } ≃o { p : Ideal A | p ∣ J } := by have f_surj : Function.Surjective (f : R ⧸ I →+* A ⧸ J) := f.surjective have fsym_surj : Function.Surjective (f.symm : A ⧸ J →+* R ⧸ I) := f.symm.surjective @@ -1101,7 +1099,6 @@ theorem idealFactorsEquivOfQuotEquiv_mem_normalizedFactors_of_mem_normalizedFact /-- The bijection between the sets of normalized factors of I and J induced by a ring isomorphism `f : R/I ≅ A/J`. -/ --- @[simps apply] -- Porting note: simpNF complains about the lemmas generated by simps def normalizedFactorsEquivOfQuotEquiv (hI : I ≠ ⊥) (hJ : J ≠ ⊥) : { L : Ideal R | L ∈ normalizedFactors I } ≃ { M : Ideal A | M ∈ normalizedFactors J } where toFun j := @@ -1379,7 +1376,6 @@ variable [NormalizationMonoid R] /-- The bijection between the (normalized) prime factors of `r` and the (normalized) prime factors of `span {r}` -/ --- @[simps] -- Porting note: simpNF complains about the lemmas generated by simps noncomputable def normalizedFactorsEquivSpanNormalizedFactors {r : R} (hr : r ≠ 0) : { d : R | d ∈ normalizedFactors r } ≃ { I : Ideal R | I ∈ normalizedFactors (Ideal.span ({r} : Set R)) } := by From 409137130c4c0e7033eea0a7e369aa8607fd0973 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Sat, 23 Nov 2024 06:55:59 +0000 Subject: [PATCH 224/829] fix(Tactic/Linter): `upstreamableDecl` should treat `structure`s like `def`s (#19360) Reported on Zulip: https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/upstreamableDecl.20linter.20bug/near/483790949 We want to treat a `structure`/`inductive` declaration in the current file just like a `def` for the `upstreamableDecl` linter: potentially place a warning on the `structure` itself, but otherwise don't place warnings on downstream uses. There's a small complication where `inductive` declarations are reported to depend on their own constructors. I couldn't find a good way to determine whether any given constant is indeed a constructor declared in some given piece of syntax, so instead we allow depending on constructors just like we allow depending on theorems. Co-authored-by: Anne Baanen --- Mathlib/Tactic/Linter/UpstreamableDecl.lean | 20 ++++++++++++----- MathlibTest/MinImports.lean | 24 +++++++++++++++++++++ 2 files changed, 39 insertions(+), 5 deletions(-) diff --git a/Mathlib/Tactic/Linter/UpstreamableDecl.lean b/Mathlib/Tactic/Linter/UpstreamableDecl.lean index 86f55dc84e8bd..59beda016bc47 100644 --- a/Mathlib/Tactic/Linter/UpstreamableDecl.lean +++ b/Mathlib/Tactic/Linter/UpstreamableDecl.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa, Anne Baanen -/ import ImportGraph.Imports -import Mathlib.Lean.Expr.Basic import Mathlib.Tactic.MinImports /-! # The `upstreamableDecl` linter @@ -21,8 +20,12 @@ def Lean.Name.isLocal (env : Environment) (decl : Name) : Bool := open Mathlib.Command.MinImports -/-- Does the declaration with this name depend on `def`s in the current file? -/ -def Lean.Environment.localDefDependencies (env : Environment) (stx id : Syntax) : +/-- Does the declaration with this name depend on definitions in the current file? + +Here, "definition" means everything that is not a theorem, and so includes `def`, +`structure`, `inductive`, etc. +-/ +def Lean.Environment.localDefinitionDependencies (env : Environment) (stx id : Syntax) : CommandElabM Bool := do let declName : NameSet ← try NameSet.ofList <$> resolveGlobalConst id @@ -37,7 +40,14 @@ def Lean.Environment.localDefDependencies (env : Environment) (stx id : Syntax) let deps ← liftCoreM <| immediateDeps.transitivelyUsedConstants let constInfos := deps.toList.filterMap env.find? - let defs := constInfos.filter (·.isDef) + -- We allow depending on theorems and constructors. + -- We explicitly allow constructors since `inductive` declarations are reported to depend on their + -- own constructors, and we want inductives to behave the same as definitions, so place one + -- warning on the inductive itself but nothing on its downstream uses. + -- (There does not seem to be an easy way to determine, given `Syntax` and `ConstInfo`, + -- whether the `ConstInfo` is a constructor declared in this piece of `Syntax`.) + let defs := constInfos.filter (fun constInfo => !(constInfo.isTheorem || constInfo.isCtor)) + return defs.any fun constInfo => !(declName.contains constInfo.name) && constInfo.name.isLocal env namespace Mathlib.Linter @@ -72,7 +82,7 @@ def upstreamableDeclLinter : Linter where run := withSetOptionIn fun stx ↦ do let minImports := getIrredundantImports env (← getAllImports stx id) match minImports with | ⟨(RBNode.node _ .leaf upstream _ .leaf), _⟩ => do - if !(← env.localDefDependencies stx id) then + if !(← env.localDefinitionDependencies stx id) then let p : GoToModuleLinkProps := { modName := upstream } let widget : MessageData := .ofWidget (← liftCoreM <| Widget.WidgetInstance.ofHash diff --git a/MathlibTest/MinImports.lean b/MathlibTest/MinImports.lean index 43ee73fa011e9..87a0a011d07ad 100644 --- a/MathlibTest/MinImports.lean +++ b/MathlibTest/MinImports.lean @@ -165,4 +165,28 @@ def def_with_multiple_dependencies := let _ := Mathlib.Meta.NormNum.evalNatDvd false +/-! Structures and inductives should be treated just like definitions. -/ + +/-- + +warning: Consider moving this declaration to the module Mathlib.Data.Nat.Notation. +note: this linter can be disabled with `set_option linter.upstreamableDecl false` +-/ +#guard_msgs in +structure ProposeToMoveThisStructure where + foo : ℕ + +/-- +warning: Consider moving this declaration to the module Mathlib.Data.Nat.Notation. +note: this linter can be disabled with `set_option linter.upstreamableDecl false` +-/ +#guard_msgs in +inductive ProposeToMoveThisInductive where +| foo : ℕ → ProposeToMoveThisInductive +| bar : ℕ → ProposeToMoveThisInductive → ProposeToMoveThisInductive + +-- This theorem depends on a local inductive, so should not be moved. +#guard_msgs in +theorem theorem_with_local_inductive : Nonempty ProposeToMoveThisInductive := ⟨.foo 0⟩ + end Linter.UpstreamableDecl From dde4596f8b160e9e0026507e6d34b40776bc2e65 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 23 Nov 2024 07:36:13 +0000 Subject: [PATCH 225/829] chore: cleanup of many set_option deprecated.linter false (#19181) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Violeta Hernández --- Mathlib.lean | 2 +- Mathlib/Algebra/AddTorsor.lean | 10 ++++++++-- Mathlib/Algebra/BigOperators/Group/List.lean | 3 ++- .../Algebra/GroupWithZero/Units/Basic.lean | 4 ++-- Mathlib/Algebra/Ring/Parity.lean | 2 -- Mathlib/Data/FP/Basic.lean | 1 - Mathlib/Data/List/Basic.lean | 12 ++++++----- Mathlib/Data/List/Range.lean | 1 + Mathlib/Data/Multiset/Basic.lean | 7 +------ Mathlib/Data/Nat/Defs.lean | 1 - Mathlib/Data/Num/Lemmas.lean | 1 - Mathlib/Data/Num/Prime.lean | 1 - .../Basic.lean => Deprecated/LazyList.lean} | 0 .../SpecificGroups/Alternating.lean | 1 - .../AffineSpace/AffineEquiv.lean | 20 +++++++++++-------- .../CliffordAlgebra/Grading.lean | 13 +++++------- .../ExteriorAlgebra/Grading.lean | 10 ++-------- Mathlib/Logic/Denumerable.lean | 1 - Mathlib/Logic/Equiv/Nat.lean | 1 - .../Measure/MeasureSpaceDef.lean | 4 ++-- Mathlib/ModelTheory/Encoding.lean | 7 ------- Mathlib/Order/OmegaCompletePartialOrder.lean | 4 ++-- Mathlib/Order/RelClasses.lean | 2 -- Mathlib/SetTheory/Cardinal/Aleph.lean | 1 + Mathlib/SetTheory/Cardinal/Cofinality.lean | 1 - scripts/noshake.json | 1 + 26 files changed, 47 insertions(+), 64 deletions(-) rename Mathlib/{Data/LazyList/Basic.lean => Deprecated/LazyList.lean} (100%) diff --git a/Mathlib.lean b/Mathlib.lean index 075b258277b1b..48be73827fdd6 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2507,7 +2507,6 @@ import Mathlib.Data.Int.Sqrt import Mathlib.Data.Int.Star import Mathlib.Data.Int.SuccPred import Mathlib.Data.Int.WithZero -import Mathlib.Data.LazyList.Basic import Mathlib.Data.List.AList import Mathlib.Data.List.Basic import Mathlib.Data.List.Chain @@ -2892,6 +2891,7 @@ import Mathlib.Deprecated.Combinator import Mathlib.Deprecated.Equiv import Mathlib.Deprecated.Group import Mathlib.Deprecated.HashMap +import Mathlib.Deprecated.LazyList import Mathlib.Deprecated.Logic import Mathlib.Deprecated.MinMax import Mathlib.Deprecated.NatLemmas diff --git a/Mathlib/Algebra/AddTorsor.lean b/Mathlib/Algebra/AddTorsor.lean index 3f6183df46993..5c7a2b005055b 100644 --- a/Mathlib/Algebra/AddTorsor.lean +++ b/Mathlib/Algebra/AddTorsor.lean @@ -404,13 +404,16 @@ theorem pointReflection_involutive (x : P) : Involutive (pointReflection x : P /-- `x` is the only fixed point of `pointReflection x`. This lemma requires `x + x = y + y ↔ x = y`. There is no typeclass to use here, so we add it as an explicit argument. -/ -theorem pointReflection_fixed_iff_of_injective_bit0 {x y : P} (h : Injective (2 • · : G → G)) : +theorem pointReflection_fixed_iff_of_injective_two_nsmul {x y : P} (h : Injective (2 • · : G → G)) : pointReflection x y = y ↔ y = x := by rw [pointReflection_apply, eq_comm, eq_vadd_iff_vsub_eq, ← neg_vsub_eq_vsub_rev, neg_eq_iff_add_eq_zero, ← two_nsmul, ← nsmul_zero 2, h.eq_iff, vsub_eq_zero_iff_eq, eq_comm] +@[deprecated (since := "2024-11-18")] alias pointReflection_fixed_iff_of_injective_bit0 := +pointReflection_fixed_iff_of_injective_two_nsmul + -- Porting note: need this to calm down CI -theorem injective_pointReflection_left_of_injective_bit0 {G P : Type*} [AddCommGroup G] +theorem injective_pointReflection_left_of_injective_two_nsmul {G P : Type*} [AddCommGroup G] [AddTorsor G P] (h : Injective (2 • · : G → G)) (y : P) : Injective fun x : P => pointReflection x y := fun x₁ x₂ (hy : pointReflection x₁ y = pointReflection x₂ y) => by @@ -418,6 +421,9 @@ theorem injective_pointReflection_left_of_injective_bit0 {G P : Type*} [AddCommG vsub_sub_vsub_cancel_right, ← neg_vsub_eq_vsub_rev, neg_eq_iff_add_eq_zero, ← two_nsmul, ← nsmul_zero 2, h.eq_iff, vsub_eq_zero_iff_eq] at hy +@[deprecated (since := "2024-11-18")] alias injective_pointReflection_left_of_injective_bit0 := +injective_pointReflection_left_of_injective_two_nsmul + end Equiv theorem AddTorsor.subsingleton_iff (G P : Type*) [AddGroup G] [AddTorsor G P] : diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index 59dc2be68d10a..bf01489af4f70 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -657,7 +657,8 @@ theorem ranges_nodup {l s : List ℕ} (hs : s ∈ ranges l) : s.Nodup := /-- Any entry of any member of `l.ranges` is strictly smaller than `l.sum`. -/ 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] + (∃ s ∈ l.ranges, n ∈ s) ↔ n < l.sum := by + rw [← mem_range, ← ranges_flatten, mem_flatten] @[simp] theorem length_flatMap (l : List α) (f : α → List β) : diff --git a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean index 4eed1a7034205..be6aebca0a776 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean @@ -422,9 +422,9 @@ instance (priority := 100) CommGroupWithZero.toDivisionCommMonoid : lemma div_mul_cancel_left₀ (ha : a ≠ 0) (b : G₀) : a / (a * b) = b⁻¹ := ha.isUnit.div_mul_cancel_left _ -set_option linter.deprecated false in @[deprecated div_mul_cancel_left₀ (since := "2024-03-22")] -lemma div_mul_right (b : G₀) (ha : a ≠ 0) : a / (a * b) = 1 / b := ha.isUnit.div_mul_right _ +lemma div_mul_right (b : G₀) (ha : a ≠ 0) : a / (a * b) = 1 / b := by + simp [div_mul_cancel_left₀ ha] lemma mul_div_cancel_left_of_imp (h : a = 0 → b = 0) : a * b / a = b := by rw [mul_comm, mul_div_cancel_of_imp h] diff --git a/Mathlib/Algebra/Ring/Parity.lean b/Mathlib/Algebra/Ring/Parity.lean index ff8df726aad0a..7b3959696657a 100644 --- a/Mathlib/Algebra/Ring/Parity.lean +++ b/Mathlib/Algebra/Ring/Parity.lean @@ -89,7 +89,6 @@ lemma Even.pow_of_ne_zero (ha : Even a) : ∀ {n : ℕ}, n ≠ 0 → Even (a ^ n /-- An element `a` of a semiring is odd if there exists `k` such `a = 2*k + 1`. -/ def Odd (a : α) : Prop := ∃ k, a = 2 * k + 1 -set_option linter.deprecated false in lemma odd_iff_exists_bit1 : Odd a ↔ ∃ b, a = 2 * b + 1 := exists_congr fun b ↦ by rw [two_mul] alias ⟨Odd.exists_bit1, _⟩ := odd_iff_exists_bit1 @@ -247,7 +246,6 @@ lemma mod_two_add_add_odd_mod_two (m : ℕ) {n : ℕ} (hn : Odd n) : m % 2 + (m lemma even_add' : Even (m + n) ↔ (Odd m ↔ Odd n) := by rw [even_add, ← not_odd_iff_even, ← not_odd_iff_even, not_iff_not] -set_option linter.deprecated false in @[simp] lemma not_even_bit1 (n : ℕ) : ¬Even (2 * n + 1) := by simp [parity_simps] lemma not_even_two_mul_add_one (n : ℕ) : ¬ Even (2 * n + 1) := diff --git a/Mathlib/Data/FP/Basic.lean b/Mathlib/Data/FP/Basic.lean index ebd1995796919..492cf4ffcd604 100644 --- a/Mathlib/Data/FP/Basic.lean +++ b/Mathlib/Data/FP/Basic.lean @@ -146,7 +146,6 @@ unsafe def nextUpPos (e m) (v : ValidFinite e m) : Float := Float.finite false e m' (by unfold ValidFinite at *; rw [ss]; exact v) else if h : e = emax then Float.inf false else Float.finite false e.succ (Nat.div2 m') lcProof -set_option linter.deprecated false in -- Porting note: remove this line when you dropped 'lcProof' set_option linter.unusedVariables false in @[nolint docBlame] diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index fac7238168886..4e3ce8533a890 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -737,7 +737,6 @@ end IndexOf /-! ### nth element -/ section deprecated -set_option linter.deprecated false @[simp] theorem getElem?_length (l : List α) : l[l.length]? = none := getElem?_len_le le_rfl @@ -754,19 +753,20 @@ theorem getElem_map_rev (f : α → β) {l} {n : Nat} {h : n < l.length} : /-- A version of `get_map` that can be used for rewriting. -/ @[deprecated getElem_map_rev (since := "2024-06-12")] theorem get_map_rev (f : α → β) {l n} : - f (get l n) = get (map f l) ⟨n.1, (l.length_map f).symm ▸ n.2⟩ := Eq.symm (get_map _) + f (get l n) = get (map f l) ⟨n.1, (l.length_map f).symm ▸ n.2⟩ := Eq.symm (getElem_map _) theorem get_length_sub_one {l : List α} (h : l.length - 1 < l.length) : l.get ⟨l.length - 1, h⟩ = l.getLast (by rintro rfl; exact Nat.lt_irrefl 0 h) := - (getLast_eq_get l _).symm + (getLast_eq_getElem l _).symm theorem take_one_drop_eq_of_lt_length {l : List α} {n : ℕ} (h : n < l.length) : (l.drop n).take 1 = [l.get ⟨n, h⟩] := by - rw [drop_eq_get_cons h, take, take] + rw [drop_eq_getElem_cons h, take, take] + simp theorem ext_get?' {l₁ l₂ : List α} (h' : ∀ n < max l₁.length l₂.length, l₁.get? n = l₂.get? n) : l₁ = l₂ := by - apply ext + apply ext_get? intro n rcases Nat.lt_or_ge n <| max l₁.length l₂.length with hn | hn · exact h' n hn @@ -831,6 +831,8 @@ theorem get_reverse (l : List α) (i : Nat) (h1 h2) : dsimp omega +set_option linter.deprecated false + theorem get_reverse' (l : List α) (n) (hn') : l.reverse.get n = l.get ⟨l.length - 1 - n, hn'⟩ := by rw [eq_comm] diff --git a/Mathlib/Data/List/Range.lean b/Mathlib/Data/List/Range.lean index 6a342b20f3573..7cefc522da8e2 100644 --- a/Mathlib/Data/List/Range.lean +++ b/Mathlib/Data/List/Range.lean @@ -149,6 +149,7 @@ lemma ranges_flatten' : ∀ l : List ℕ, l.ranges.flatten = range (Nat.sum l) set_option linter.deprecated false in /-- Any entry of any member of `l.ranges` is strictly smaller than `Nat.sum l`. See `List.mem_mem_ranges_iff_lt_sum` for the version about `List.sum`. -/ +@[deprecated "Use `List.mem_mem_ranges_iff_lt_sum`." (since := "2024-11-18")] lemma mem_mem_ranges_iff_lt_natSum (l : List ℕ) {n : ℕ} : (∃ s ∈ l.ranges, n ∈ s) ↔ n < Nat.sum l := by rw [← mem_range, ← ranges_flatten', mem_flatten] diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index 49f764b115ba5..e364da53cefec 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -1720,15 +1720,10 @@ def filter (s : Multiset α) : Multiset α := theorem filter_zero : filter p 0 = 0 := rfl -#adaptation_note -/-- -Please re-enable the linter once we moved to `nightly-2024-06-22` or later. --/ -set_option linter.deprecated false in @[congr] theorem filter_congr {p q : α → Prop} [DecidablePred p] [DecidablePred q] {s : Multiset α} : (∀ x ∈ s, p x ↔ q x) → filter p s = filter q s := - Quot.inductionOn s fun _l h => congr_arg ofList <| filter_congr' <| by simpa using h + Quot.inductionOn s fun _l h => congr_arg ofList <| List.filter_congr <| by simpa using h @[simp] theorem filter_add (s t : Multiset α) : filter p (s + t) = filter p s + filter p t := diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index 2327e2c3bb37a..85ee5070c28b1 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -1168,7 +1168,6 @@ lemma mul_add_mod' (a b c : ℕ) : (a * b + c) % b = c % b := by rw [Nat.mul_com lemma mul_add_mod_of_lt (h : c < b) : (a * b + c) % b = c := by rw [Nat.mul_add_mod', Nat.mod_eq_of_lt h] -set_option linter.deprecated false in @[simp] protected theorem not_two_dvd_bit1 (n : ℕ) : ¬2 ∣ 2 * n + 1 := by omega diff --git a/Mathlib/Data/Num/Lemmas.lean b/Mathlib/Data/Num/Lemmas.lean index 04f10c375c5b9..c7e98484c0e1c 100644 --- a/Mathlib/Data/Num/Lemmas.lean +++ b/Mathlib/Data/Num/Lemmas.lean @@ -308,7 +308,6 @@ theorem of_to_nat' : ∀ n : Num, Num.ofNat' (n : ℕ) = n | 0 => ofNat'_zero | pos p => p.of_to_nat' -set_option linter.deprecated false in lemma toNat_injective : Injective (castNum : Num → ℕ) := LeftInverse.injective of_to_nat' @[norm_cast] diff --git a/Mathlib/Data/Num/Prime.lean b/Mathlib/Data/Num/Prime.lean index 6b3a141d15139..c183f23216e12 100644 --- a/Mathlib/Data/Num/Prime.lean +++ b/Mathlib/Data/Num/Prime.lean @@ -37,7 +37,6 @@ def minFacAux (n : PosNum) : ℕ → PosNum → PosNum | fuel + 1, k => if n < k.bit1 * k.bit1 then n else if k.bit1 ∣ n then k.bit1 else minFacAux n fuel k.succ -set_option linter.deprecated false in theorem minFacAux_to_nat {fuel : ℕ} {n k : PosNum} (h : Nat.sqrt n < fuel + k.bit1) : (minFacAux n fuel k : ℕ) = Nat.minFacAux n k.bit1 := by induction' fuel with fuel ih generalizing k <;> rw [minFacAux, Nat.minFacAux] diff --git a/Mathlib/Data/LazyList/Basic.lean b/Mathlib/Deprecated/LazyList.lean similarity index 100% rename from Mathlib/Data/LazyList/Basic.lean rename to Mathlib/Deprecated/LazyList.lean diff --git a/Mathlib/GroupTheory/SpecificGroups/Alternating.lean b/Mathlib/GroupTheory/SpecificGroups/Alternating.lean index 982ee49d373ef..8b0f82d484a0d 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Alternating.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Alternating.lean @@ -78,7 +78,6 @@ theorem IsThreeCycle.mem_alternatingGroup {f : Perm α} (h : IsThreeCycle f) : f ∈ alternatingGroup α := Perm.mem_alternatingGroup.mpr h.sign -set_option linter.deprecated false in theorem finRotate_bit1_mem_alternatingGroup {n : ℕ} : finRotate (2 * n + 1) ∈ alternatingGroup (Fin (2 * n + 1)) := by rw [mem_alternatingGroup, sign_finRotate, pow_mul, pow_two, Int.units_mul_self, one_pow] diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean index 63e23ebbb7c77..75e11673b7d74 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean @@ -486,22 +486,26 @@ theorem pointReflection_self (x : P₁) : pointReflection k x x = x := theorem pointReflection_involutive (x : P₁) : Involutive (pointReflection k x : P₁ → P₁) := Equiv.pointReflection_involutive x -set_option linter.deprecated false in /-- `x` is the only fixed point of `pointReflection x`. This lemma requires `x + x = y + y ↔ x = y`. There is no typeclass to use here, so we add it as an explicit argument. -/ -theorem pointReflection_fixed_iff_of_injective_bit0 {x y : P₁} (h : Injective (2 • · : V₁ → V₁)) : - pointReflection k x y = y ↔ y = x := - Equiv.pointReflection_fixed_iff_of_injective_bit0 h +theorem pointReflection_fixed_iff_of_injective_two_nsmul {x y : P₁} + (h : Injective (2 • · : V₁ → V₁)) : pointReflection k x y = y ↔ y = x := + Equiv.pointReflection_fixed_iff_of_injective_two_nsmul h + +@[deprecated (since := "2024-11-18")] alias pointReflection_fixed_iff_of_injective_bit0 := +pointReflection_fixed_iff_of_injective_two_nsmul -set_option linter.deprecated false in -theorem injective_pointReflection_left_of_injective_bit0 +theorem injective_pointReflection_left_of_injective_two_nsmul (h : Injective (2 • · : V₁ → V₁)) (y : P₁) : Injective fun x : P₁ => pointReflection k x y := - Equiv.injective_pointReflection_left_of_injective_bit0 h y + Equiv.injective_pointReflection_left_of_injective_two_nsmul h y + +@[deprecated (since := "2024-11-18")] alias injective_pointReflection_left_of_injective_bit0 := +injective_pointReflection_left_of_injective_two_nsmul theorem injective_pointReflection_left_of_module [Invertible (2 : k)] : ∀ y, Injective fun x : P₁ => pointReflection k x y := - injective_pointReflection_left_of_injective_bit0 k fun x y h => by + injective_pointReflection_left_of_injective_two_nsmul k fun x y h => by dsimp at h rwa [two_nsmul, two_nsmul, ← two_smul k x, ← two_smul k y, (isUnit_of_invertible (2 : k)).smul_left_cancel] at h diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean index 39224b85617fe..d88ad52680183 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean @@ -75,7 +75,6 @@ nonrec theorem GradedAlgebra.ι_sq_scalar (m : M) : rw [GradedAlgebra.ι_apply Q, DirectSum.of_mul_of, DirectSum.algebraMap_apply] exact DirectSum.of_eq_of_gradedMonoid_eq (Sigma.subtype_ext rfl <| ι_sq_scalar _ _) -set_option linter.deprecated false in theorem GradedAlgebra.lift_ι_eq (i' : ZMod 2) (x' : evenOdd Q i') : -- Porting note: added a second `by apply` lift Q ⟨by apply GradedAlgebra.ι Q, by apply GradedAlgebra.ι_sq_scalar Q⟩ x' = @@ -85,29 +84,27 @@ theorem GradedAlgebra.lift_ι_eq (i' : ZMod 2) (x' : evenOdd Q i') : induction hx' using Submodule.iSup_induction' with | mem i x hx => obtain ⟨i, rfl⟩ := i - -- Porting note: `dsimp only [Subtype.coe_mk] at hx` doesn't work, use `change` instead - change x ∈ LinearMap.range (ι Q) ^ i at hx + dsimp only [Subtype.coe_mk] at hx induction hx using Submodule.pow_induction_on_left' with | algebraMap r => rw [AlgHom.commutes, DirectSum.algebraMap_apply]; rfl | add x y i hx hy ihx ihy => - -- Note: in https://github.com/leanprover-community/mathlib4/pull/8386 `map_add` had to be specialized to avoid a timeout - -- (the definition was already very slow) - rw [AlgHom.map_add, ihx, ihy, ← AddMonoidHom.map_add] + rw [map_add, ihx, ihy, ← AddMonoidHom.map_add] rfl | mem_mul m hm i x hx ih => obtain ⟨_, rfl⟩ := hm - rw [AlgHom.map_mul, ih, lift_ι_apply, GradedAlgebra.ι_apply Q, DirectSum.of_mul_of] + rw [map_mul, ih, lift_ι_apply, GradedAlgebra.ι_apply Q, DirectSum.of_mul_of] refine DirectSum.of_eq_of_gradedMonoid_eq (Sigma.subtype_ext ?_ ?_) <;> dsimp only [GradedMonoid.mk, Subtype.coe_mk] · rw [Nat.succ_eq_add_one, add_comm, Nat.cast_add, Nat.cast_one] rfl | zero => + set_option linter.deprecated false in rw [AlgHom.map_zero] apply Eq.symm apply DFinsupp.single_eq_zero.mpr; rfl | add x y hx hy ihx ihy => - rw [AlgHom.map_add, ihx, ihy, ← AddMonoidHom.map_add]; rfl + rw [map_add, ihx, ihy, ← AddMonoidHom.map_add]; rfl /-- The clifford algebra is graded by the even and odd parts. -/ instance gradedAlgebra : GradedAlgebra (evenOdd Q) := diff --git a/Mathlib/LinearAlgebra/ExteriorAlgebra/Grading.lean b/Mathlib/LinearAlgebra/ExteriorAlgebra/Grading.lean index 81a29f1543817..d39a95c8a5d29 100644 --- a/Mathlib/LinearAlgebra/ExteriorAlgebra/Grading.lean +++ b/Mathlib/LinearAlgebra/ExteriorAlgebra/Grading.lean @@ -54,22 +54,16 @@ def GradedAlgebra.liftι : ExteriorAlgebra R M →ₐ[R] ⨁ i : ℕ, ⋀[R]^i M := lift R ⟨by apply GradedAlgebra.ι R M, GradedAlgebra.ι_sq_zero R M⟩ -set_option linter.deprecated false in theorem GradedAlgebra.liftι_eq (i : ℕ) (x : ⋀[R]^i M) : GradedAlgebra.liftι R M x = DirectSum.of (fun i => ⋀[R]^i M) i x := by cases' x with x hx dsimp only [Subtype.coe_mk, DirectSum.lof_eq_of] - -- Porting note: original statement was - -- refine Submodule.pow_induction_on_left' _ (fun r => ?_) (fun x y i hx hy ihx ihy => ?_) - -- (fun m hm i x hx ih => ?_) hx - -- but it created invalid goals induction hx using Submodule.pow_induction_on_left' with | algebraMap => simp_rw [AlgHom.commutes, DirectSum.algebraMap_apply]; rfl - -- FIXME: specialized `map_add` to avoid a (whole-declaration) timeout - | add _ _ _ _ _ ihx ihy => simp_rw [AlgHom.map_add, ihx, ihy, ← AddMonoidHom.map_add]; rfl + | add _ _ _ _ _ ihx ihy => simp_rw [map_add, ihx, ihy, ← AddMonoidHom.map_add]; rfl | mem_mul _ hm _ _ _ ih => obtain ⟨_, rfl⟩ := hm - simp_rw [AlgHom.map_mul, ih, GradedAlgebra.liftι, lift_ι_apply, GradedAlgebra.ι_apply R M, + simp_rw [map_mul, ih, GradedAlgebra.liftι, lift_ι_apply, GradedAlgebra.ι_apply R M, DirectSum.of_mul_of] exact DirectSum.of_eq_of_gradedMonoid_eq (Sigma.subtype_ext (add_comm _ _) rfl) diff --git a/Mathlib/Logic/Denumerable.lean b/Mathlib/Logic/Denumerable.lean index f96dccd8f839c..d6160e1858ab6 100644 --- a/Mathlib/Logic/Denumerable.lean +++ b/Mathlib/Logic/Denumerable.lean @@ -121,7 +121,6 @@ instance option : Denumerable (Option α) := · rw [decode_option_succ, decode_eq_ofNat, Option.map_some', Option.mem_def] rw [encode_some, encode_ofNat]⟩ -set_option linter.deprecated false in /-- If `α` and `β` are denumerable, then so is their sum. -/ instance sum : Denumerable (α ⊕ β) := ⟨fun n => by diff --git a/Mathlib/Logic/Equiv/Nat.lean b/Mathlib/Logic/Equiv/Nat.lean index 0bc17024d2609..69466efcdac2d 100644 --- a/Mathlib/Logic/Equiv/Nat.lean +++ b/Mathlib/Logic/Equiv/Nat.lean @@ -36,7 +36,6 @@ def boolProdNatEquivNat : Bool × ℕ ≃ ℕ where def natSumNatEquivNat : ℕ ⊕ ℕ ≃ ℕ := (boolProdEquivSum ℕ).symm.trans boolProdNatEquivNat -set_option linter.deprecated false in @[simp] theorem natSumNatEquivNat_apply : ⇑natSumNatEquivNat = Sum.elim (2 * ·) (2 * · + 1) := by ext (x | x) <;> rfl diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean b/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean index 8eb64f43a20e2..11481cc61993f 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean @@ -82,9 +82,9 @@ instance Measure.instFunLike [MeasurableSpace α] : FunLike (Measure α) (Set α coe μ := μ.toOuterMeasure coe_injective' | ⟨_, _, _⟩, ⟨_, _, _⟩, h => toOuterMeasure_injective <| DFunLike.coe_injective h -set_option linter.deprecated false in -- Not immediately obvious how to use `measure_empty` here. + instance Measure.instOuterMeasureClass [MeasurableSpace α] : OuterMeasureClass (Measure α) α where - measure_empty m := m.empty' + measure_empty m := measure_empty (μ := m.toOuterMeasure) measure_iUnion_nat_le m := m.iUnion_nat measure_mono m := m.mono diff --git a/Mathlib/ModelTheory/Encoding.lean b/Mathlib/ModelTheory/Encoding.lean index 5395d5324a5a6..180cb86b870ed 100644 --- a/Mathlib/ModelTheory/Encoding.lean +++ b/Mathlib/ModelTheory/Encoding.lean @@ -181,13 +181,6 @@ or returns `default` if not possible. -/ def sigmaImp : (Σn, L.BoundedFormula α n) → (Σn, L.BoundedFormula α n) → Σn, L.BoundedFormula α n | ⟨m, φ⟩, ⟨n, ψ⟩ => if h : m = n then ⟨m, φ.imp (Eq.mp (by rw [h]) ψ)⟩ else default -#adaptation_note -/-- -`List.drop_sizeOf_le` is deprecated. -See https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Why.20is.20.60Mathlib.2EModelTheory.2EEncoding.60.20using.20.60SizeOf.2EsizeOf.60.3F -for discussion about adapting this code. --/ -set_option linter.deprecated false in /-- Decodes a list of symbols as a list of formulas. -/ @[simp] lemma sigmaImp_apply {n} {φ ψ : L.BoundedFormula α n} : diff --git a/Mathlib/Order/OmegaCompletePartialOrder.lean b/Mathlib/Order/OmegaCompletePartialOrder.lean index 309394381dfc2..4c97970762171 100644 --- a/Mathlib/Order/OmegaCompletePartialOrder.lean +++ b/Mathlib/Order/OmegaCompletePartialOrder.lean @@ -802,11 +802,11 @@ theorem seq_continuous' {β γ : Type v} (f : α → Part (β → γ)) (g : α intro apply map_continuous' _ _ hg +set_option linter.deprecated true + theorem continuous (F : α →𝒄 β) (C : Chain α) : F (ωSup C) = ωSup (C.map F) := F.ωScottContinuous.map_ωSup _ -set_option linter.deprecated true - /-- Construct a continuous function from a bare function, a continuous function, and a proof that they are equal. -/ -- Porting note: removed `@[reducible]` diff --git a/Mathlib/Order/RelClasses.lean b/Mathlib/Order/RelClasses.lean index 6fb587d189e2c..c99abc227fd09 100644 --- a/Mathlib/Order/RelClasses.lean +++ b/Mathlib/Order/RelClasses.lean @@ -17,8 +17,6 @@ extend `LE` and/or `LT` while these classes take a relation as an explicit argum -/ -set_option linter.deprecated false - universe u v variable {α : Type u} {β : Type v} {r : α → α → Prop} {s : β → β → Prop} diff --git a/Mathlib/SetTheory/Cardinal/Aleph.lean b/Mathlib/SetTheory/Cardinal/Aleph.lean index 3970c6fc6cb57..dafc9ee33dced 100644 --- a/Mathlib/SetTheory/Cardinal/Aleph.lean +++ b/Mathlib/SetTheory/Cardinal/Aleph.lean @@ -616,6 +616,7 @@ alias aleph'_omega := aleph'_omega0 def aleph'Equiv : Ordinal ≃ Cardinal := ⟨aleph', alephIdx, alephIdx_aleph', aleph'_alephIdx⟩ +@[deprecated aleph_eq_preAleph (since := "2024-10-22")] theorem aleph_eq_aleph' (o : Ordinal) : ℵ_ o = preAleph (ω + o) := rfl diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index 4024caa0c1f7a..87b3654599332 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -287,7 +287,6 @@ theorem lsub_lt_ord {ι} {f : ι → Ordinal} {c : Ordinal} (hι : #ι < c.cof) (∀ i, f i < c) → lsub.{u, u} f < c := lsub_lt_ord_lift (by rwa [(#ι).lift_id]) -set_option linter.deprecated false in theorem cof_iSup_le_lift {ι} {f : ι → Ordinal} (H : ∀ i, f i < iSup f) : cof (iSup f) ≤ Cardinal.lift.{v, u} #ι := by rw [← Ordinal.sup] at * diff --git a/scripts/noshake.json b/scripts/noshake.json index 736e94310eaf9..86cfab061a740 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -364,6 +364,7 @@ "Mathlib.Util.Superscript"], "Mathlib.Deprecated.NatLemmas": ["Batteries.Data.Nat.Lemmas", "Batteries.WF"], "Mathlib.Deprecated.MinMax": ["Mathlib.Order.MinMax"], + "Mathlib.Deprecated.LazyList": ["Mathlib.Lean.Thunk"], "Mathlib.Deprecated.ByteArray": ["Batteries.Data.ByteSubarray"], "Mathlib.Data.Vector.Basic": ["Mathlib.Control.Applicative"], "Mathlib.Data.Set.Image": From 84d03da261895eac598b0ac33120abfe8ee711f0 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Sat, 23 Nov 2024 08:16:21 +0000 Subject: [PATCH 226/829] feat(LinearAlgebra/QuadraticForm/Basic): weaken invertibility hypothesis on 2 (#14986) The aim of this is to eventually enable base-change of quadratic maps from a (not necessarily (torsion-)free) abelian group to an abelian group on which 2 is invertible (e.g., the real numbers). The current set-up requires 2 to be invertible in the base ring (which would be the integers in the situation above). This PR weakens the assumption for the definition of `QuadraticMap.associatedHom` and `QuadraticMap.associated` from `[Invertible (2 : R)]` to `[Invertible (2 : Module.End R N)]`, where `N` is the target module of he quadratic map. See [this discussion](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/Quadratic.20forms.20with.20values.20in.20a.20larger.20ring/near/450582700) on Zulip. On the way, we fix some mistakes in docstrings. --- .../LinearAlgebra/QuadraticForm/Basic.lean | 163 ++++++++++++------ Mathlib/LinearAlgebra/QuadraticForm/Dual.lean | 7 +- 2 files changed, 119 insertions(+), 51 deletions(-) diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean index bec420998fa64..9778e86170c13 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean @@ -526,8 +526,8 @@ variable [CommSemiring R] [CommSemiring S] [AddCommMonoid M] [Module R M] [AddCo [Module R N] [Module S M] [Module S N] [Algebra S R] variable [IsScalarTower S R M] [IsScalarTower S R N] -/-- If `B : M → N → Pₗ` is `R`-`S` bilinear and `R'` and `S'` are compatible scalar multiplications, -then the restriction of scalars is a `R'`-`S'` bilinear map. -/ +/-- If `Q : M → N` is a quadratic map of `R`-modules and `R` is an `S`-algebra, +then the restriction of scalars is a quadratic map of `S`-modules. -/ @[simps!] def restrictScalars (Q : QuadraticMap R M N) : QuadraticMap S M N where toFun x := Q x @@ -597,12 +597,13 @@ theorem _root_.LinearEquiv.congrQuadraticMap_symm (e : N ≃ₗ[R] P) : (LinearEquiv.congrQuadraticMap e (M := M)).symm = e.symm.congrQuadraticMap := rfl end Comp + section NonUnitalNonAssocSemiring variable [CommSemiring R] [NonUnitalNonAssocSemiring A] [AddCommMonoid M] [Module R M] variable [Module R A] [SMulCommClass R A A] [IsScalarTower R A A] -/-- The product of linear forms is a quadratic form. -/ +/-- The product of linear maps into an `R`-algebra is a quadratic map. -/ def linMulLin (f g : M →ₗ[R] A) : QuadraticMap R M A where toFun := f * g toFun_smul a x := by @@ -635,12 +636,12 @@ theorem linMulLin_comp (f g : M →ₗ[R] A) (h : N' →ₗ[R] M) : variable {n : Type*} -/-- `sq` is the quadratic form mapping the vector `x : A` to `x * x` -/ +/-- `sq` is the quadratic map sending the vector `x : A` to `x * x` -/ @[simps!] def sq : QuadraticMap R A A := linMulLin LinearMap.id LinearMap.id -/-- `proj i j` is the quadratic form mapping the vector `x : n → R` to `x i * x j` -/ +/-- `proj i j` is the quadratic map sending the vector `x : n → R` to `x i * x j` -/ def proj (i j : n) : QuadraticMap R (n → A) A := linMulLin (@LinearMap.proj _ _ _ (fun _ => A) _ _ i) (@LinearMap.proj _ _ _ (fun _ => A) _ _ j) @@ -655,10 +656,16 @@ end QuadraticMap /-! ### Associated bilinear maps -Over a commutative ring with an inverse of 2, the theory of quadratic maps is -basically identical to that of symmetric bilinear maps. The map from quadratic -maps to bilinear maps giving this identification is called the `QuadraticMap.associated` -quadratic map. +If multiplication by 2 is invertible on the target module `N` of +`QuadraticMap R M N`, then there is a linear bijection `QuadraticMap.associated` +between quadratic maps `Q` over `R` from `M` to `N` and symmetric bilinear maps +`B : M →ₗ[R] M →ₗ[R] → N` such that `BilinMap.toQuadraticMap B = Q` +(see `QuadraticMap.associated_rightInverse`). The associated bilinear map is half +`Q.polarBilin` (see `QuadraticMap.two_nsmul_associated`); this is where the invertibility condition +comes from. We spell the condition as `[Invertible (2 : Module.End R N)]`. + +Note that this makes the bijection available in more cases than the simpler condition +`Invertible (2 : R)`, e.g., when `R = ℤ` and `N = ℝ`. -/ namespace LinearMap @@ -711,14 +718,14 @@ section variable (S R M) -/-- `LinearMap.BilinForm.toQuadraticMap` as an additive homomorphism -/ +/-- `LinearMap.BilinMap.toQuadraticMap` as an additive homomorphism -/ @[simps] def toQuadraticMapAddMonoidHom : (BilinMap R M N) →+ QuadraticMap R M N where toFun := toQuadraticMap map_zero' := toQuadraticMap_zero _ _ map_add' := toQuadraticMap_add -/-- `LinearMap.BilinForm.toQuadraticMap` as a linear map -/ +/-- `LinearMap.BilinMap.toQuadraticMap` as a linear map -/ @[simps!] def toQuadraticMapLinearMap [Semiring S] [Module S N] [SMulCommClass S R N] [SMulCommClass R S N] : (BilinMap R M N) →ₗ[S] QuadraticMap R M N where @@ -818,43 +825,87 @@ namespace QuadraticMap open LinearMap (BilinMap) +section + +variable [Semiring R] [AddCommMonoid M] [Module R M] + +instance : SMulCommClass R (Submonoid.center R) M where + smul_comm r r' m := by + simp_rw [Submonoid.smul_def, smul_smul, (Set.mem_center_iff.1 r'.prop).1] + +/-- If `2` is invertible in `R`, then it is also invertible in `End R M`. -/ +instance [Invertible (2 : R)] : Invertible (2 : Module.End R M) where + invOf := (⟨⅟2, Set.invOf_mem_center (Set.ofNat_mem_center _ _)⟩ : Submonoid.center R) • + (1 : Module.End R M) + invOf_mul_self := by + ext m + dsimp [Submonoid.smul_def] + rw [← ofNat_smul_eq_nsmul R, invOf_smul_smul (2 : R) m] + mul_invOf_self := by + ext m + dsimp [Submonoid.smul_def] + rw [← ofNat_smul_eq_nsmul R, smul_invOf_smul (2 : R) m] + +/-- If `2` is invertible in `R`, then applying the inverse of `2` in `End R M` to an element +of `M` is the same as multiplying by the inverse of `2` in `R`. -/ +@[simp] +lemma half_moduleEnd_apply_eq_half_smul [Invertible (2 : R)] (x : M) : + ⅟ (2 : Module.End R M) x = ⅟ (2 : R) • x := + rfl + +end + section AssociatedHom -variable [CommRing R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] -variable (S) [CommSemiring S] [Algebra S R] -variable [Module S N] [IsScalarTower S R N] -variable [Invertible (2 : R)] {B₁ : BilinMap R M R} +variable [CommRing R] [AddCommGroup M] [Module R M] +variable [AddCommGroup N] [Module R N] +variable (S) [CommSemiring S] [Algebra S R] [Module S N] [IsScalarTower S R N] + +-- the requirement that multiplication by `2` is invertible on the target module `N` +variable [Invertible (2 : Module.End R N)] /-- `associatedHom` is the map that sends a quadratic map on a module `M` over `R` to its associated symmetric bilinear map. As provided here, this has the structure of an `S`-linear map -where `S` is a commutative subring of `R`. +where `S` is a commutative ring and `R` is an `S`-algebra. Over a commutative ring, use `QuadraticMap.associated`, which gives an `R`-linear map. Over a general ring with no nontrivial distinguished commutative subring, use `QuadraticMap.associated'`, which gives an additive homomorphism (or more precisely a `ℤ`-linear map.) -/ -def associatedHom : QuadraticMap R M N →ₗ[S] (BilinMap R M N) := - -- TODO: this `center` stuff is vertigial from an incorrect non-commutative version, but we leave - -- it behind to make a future refactor to a *correct* non-commutative version easier in future. - (⟨⅟2, Set.invOf_mem_center (Set.ofNat_mem_center _ _)⟩ : Submonoid.center R) • - { toFun := polarBilin - map_add' := fun _x _y => LinearMap.ext₂ <| polar_add _ _ - map_smul' := fun _c _x => LinearMap.ext₂ <| polar_smul _ _ } +def associatedHom : QuadraticMap R M N →ₗ[S] (BilinMap R M N) where + toFun Q := ⅟ (2 : Module.End R N) • polarBilin Q + map_add' _ _ := + LinearMap.ext₂ fun _ _ ↦ by + simp only [LinearMap.smul_apply, polarBilin_apply_apply, coeFn_add, polar_add, + LinearMap.smul_def, LinearMap.map_add, LinearMap.add_apply] + map_smul' _ _ := + LinearMap.ext₂ fun _ _ ↦ by + simp only [LinearMap.smul_apply, polarBilin_apply_apply, coeFn_smul, polar_smul, + LinearMap.smul_def, LinearMap.map_smul_of_tower, RingHom.id_apply] variable (Q : QuadraticMap R M N) @[simp] -theorem associated_apply (x y : M) : associatedHom S Q x y = ⅟ (2 : R) • (Q (x + y) - Q x - Q y) := +theorem associated_apply (x y : M) : + associatedHom S Q x y = ⅟ (2 : Module.End R N) • (Q (x + y) - Q x - Q y) := rfl +/-- Twice the associated bilinear map of `Q` is the same as the polar of `Q`. -/ @[simp] theorem two_nsmul_associated : 2 • associatedHom S Q = Q.polarBilin := by ext dsimp - rw [← smul_assoc, two_nsmul, invOf_two_add_invOf_two, one_smul, polar] + rw [← LinearMap.smul_apply, nsmul_eq_mul, Nat.cast_ofNat, mul_invOf_self', LinearMap.one_apply, + polar] -theorem associated_isSymm (Q : QuadraticMap R M R) : (associatedHom S Q).IsSymm := fun x y => by - simp only [associated_apply, sub_eq_add_neg, add_assoc, map_mul, RingHom.id_apply, map_add, - _root_.map_neg, add_comm, add_left_comm] +theorem associated_isSymm (Q : QuadraticForm R M) [Invertible (2 : R)] : + (associatedHom S Q).IsSymm := fun x y ↦ by + simp only [associated_apply, sub_eq_add_neg, add_assoc, RingHom.id_apply, add_comm, add_left_comm] +/-- A version of `QuadraticMap.associated_isSymm` for general targets +(using `flip` because `IsSymm` does not apply here). -/ +lemma associated_flip : (associatedHom S Q).flip = associatedHom S Q := by + ext + simp only [LinearMap.flip_apply, associated_apply, add_comm, sub_eq_add_neg, add_left_comm, + add_assoc] @[simp] theorem associated_comp {N' : Type*} [AddCommGroup N'] [Module R N'] (f : N' →ₗ[R] M) : @@ -862,23 +913,29 @@ theorem associated_comp {N' : Type*} [AddCommGroup N'] [Module R N'] (f : N' → ext simp only [associated_apply, comp_apply, map_add, LinearMap.compl₁₂_apply] -theorem associated_toQuadraticMap (B : BilinMap R M R) (x y : M) : - associatedHom S B.toQuadraticMap x y = ⅟ (2 : R) • (B x y + B y x) := by - simp only [associated_apply, LinearMap.BilinMap.toQuadraticMap_apply, map_add, - LinearMap.add_apply, smul_eq_mul] +theorem associated_toQuadraticMap (B : BilinMap R M N) (x y : M) : + associatedHom S B.toQuadraticMap x y = ⅟ (2 : Module.End R N) • (B x y + B y x) := by + simp only [associated_apply, BilinMap.toQuadraticMap_apply, map_add, LinearMap.add_apply, + LinearMap.smul_def, _root_.map_sub] abel_nf -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] - match_scalars - linear_combination invOf_mul_self' (2:R) +theorem associated_left_inverse [Invertible (2 : R)] {B₁ : BilinMap R M R} (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_eq_mul, + invOf_smul_eq_iff, two_smul, two_smul] + +/-- A version of `QuadraticMap.associated_left_inverse` for general targets. -/ +lemma associated_left_inverse' {B₁ : BilinMap R M N} (hB₁ : B₁.flip = B₁) : + associatedHom S B₁.toQuadraticMap = B₁ := by + ext _ y + rw [associated_toQuadraticMap, ← LinearMap.flip_apply _ y, hB₁, invOf_smul_eq_iff, two_smul] -- 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] - match_scalars - linear_combination invOf_mul_self' (2:R) + 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, + invOf_smul_eq_iff, two_smul, two_smul] theorem toQuadraticMap_associated : (associatedHom S Q).toQuadraticMap = Q := QuadraticMap.ext <| associated_eq_self_apply S Q @@ -887,7 +944,7 @@ theorem toQuadraticMap_associated : (associatedHom S Q).toQuadraticMap = Q := -- with historical naming in this file. theorem associated_rightInverse : Function.RightInverse (associatedHom S) (BilinMap.toQuadraticMap : _ → QuadraticMap R M N) := - fun Q => toQuadraticMap_associated S Q + toQuadraticMap_associated S /-- `associated'` is the `ℤ`-linear map that sends a quadratic form on a module `M` over `R` to its associated symmetric bilinear form. -/ @@ -895,10 +952,15 @@ abbrev associated' : QuadraticMap R M N →ₗ[ℤ] BilinMap R M N := associatedHom ℤ /-- Symmetric bilinear forms can be lifted to quadratic forms -/ -instance canLift : +instance canLift [Invertible (2 : R)] : CanLift (BilinMap R M R) (QuadraticForm R M) (associatedHom ℕ) LinearMap.IsSymm where prf B hB := ⟨B.toQuadraticMap, associated_left_inverse _ hB⟩ +/-- Symmetric bilinear maps can be lifted to quadratic maps -/ +instance canLift' : + CanLift (BilinMap R M N) (QuadraticMap R M N) (associatedHom ℕ) fun B ↦ B.flip = B where + prf B hB := ⟨B.toQuadraticMap, associated_left_inverse' _ hB⟩ + /-- There exists a non-null vector with respect to any quadratic form `Q` whose associated bilinear form is non-zero, i.e. there exists `x` such that `Q x ≠ 0`. -/ theorem exists_quadraticMap_ne_zero {Q : QuadraticMap R M N} @@ -919,10 +981,11 @@ section Associated variable [CommSemiring S] [CommRing R] [AddCommGroup M] [Algebra S R] [Module R M] variable [AddCommGroup N] [Module R N] [Module S N] [IsScalarTower S R N] -variable [Invertible (2 : R)] +variable [Invertible (2 : Module.End R N)] -- Note: When possible, rather than writing lemmas about `associated`, write a lemma applying to -- the more general `associatedHom` and place it in the previous section. + /-- `associated` is the linear map that sends a quadratic map over a commutative ring to its associated symmetric bilinear map. -/ abbrev associated : QuadraticMap R M N →ₗ[R] BilinMap R M N := @@ -935,17 +998,19 @@ theorem coe_associatedHom : open LinearMap in @[simp] -theorem associated_linMulLin (f g : M →ₗ[R] R) : - associated (R := R) (linMulLin f g) = +theorem associated_linMulLin [Invertible (2 : R)] (f g : M →ₗ[R] R) : + associated (R := R) (N := R) (linMulLin f g) = ⅟ (2 : R) • ((mul R R).compl₁₂ f g + (mul R R).compl₁₂ g f) := by ext simp only [associated_apply, linMulLin_apply, map_add, smul_add, LinearMap.add_apply, - LinearMap.smul_apply, compl₁₂_apply, mul_apply', smul_eq_mul] + LinearMap.smul_apply, compl₁₂_apply, mul_apply', smul_eq_mul, invOf_smul_eq_iff] + simp only [smul_add, LinearMap.smul_def, Module.End.ofNat_apply, nsmul_eq_mul, Nat.cast_ofNat, + mul_invOf_cancel_left'] ring_nf open LinearMap in @[simp] -lemma associated_sq : associated (R := R) sq = mul R R := +lemma associated_sq [Invertible (2 : R)] : associated (R := R) sq = mul R R := (associated_linMulLin (id) (id)).trans <| by simp only [smul_add, invOf_two_smul_add_invOf_two_smul]; rfl @@ -1296,7 +1361,7 @@ theorem weightedSumSquares_apply [Monoid S] [DistribMulAction S R] [SMulCommClas /-- On an orthogonal basis, the basis representation of `Q` is just a sum of squares. -/ theorem basisRepr_eq_of_iIsOrtho {R M} [CommRing R] [AddCommGroup M] [Module R M] - [Invertible (2 : R)] (Q : QuadraticMap R M R) (v : Basis ι R M) + [Invertible (2 : R)] (Q : QuadraticForm R M) (v : Basis ι R M) (hv₂ : (associated (R := R) Q).IsOrthoᵢ v) : Q.basisRepr v = weightedSumSquares _ fun i => Q (v i) := by ext w @@ -1304,7 +1369,7 @@ theorem basisRepr_eq_of_iIsOrtho {R M} [CommRing R] [AddCommGroup M] [Module R M refine sum_congr rfl fun j hj => ?_ rw [← @associated_eq_self_apply R, LinearMap.map_sum₂, sum_eq_single_of_mem j hj] · rw [LinearMap.map_smul, LinearMap.map_smul₂, smul_eq_mul, associated_apply, smul_eq_mul, - smul_eq_mul, smul_eq_mul] + smul_eq_mul, LinearMap.smul_def, half_moduleEnd_apply_eq_half_smul] ring_nf · intro i _ hij rw [LinearMap.map_smul, LinearMap.map_smul₂, hv₂ hij] diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Dual.lean b/Mathlib/LinearAlgebra/QuadraticForm/Dual.lean index 48db5f20ea5c2..803a8340d71e3 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Dual.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Dual.lean @@ -141,8 +141,11 @@ def toDualProd (Q : QuadraticForm R M) [Invertible (2 : R)] : LinearMap.coe_toAddHom, LinearMap.prod_apply, Pi.prod, LinearMap.add_apply, LinearMap.coe_comp, Function.comp_apply, LinearMap.fst_apply, LinearMap.snd_apply, LinearMap.sub_apply, dualProd_apply, polarBilin_apply_apply, prod_apply, neg_apply] - simp [polar_comm _ x.1 x.2, ← sub_add, mul_sub, sub_mul, smul_sub, Submonoid.smul_def, ← - sub_eq_add_neg (Q x.1) (Q x.2)] + simp only [polar_sub_right, polar_self, nsmul_eq_mul, Nat.cast_ofNat, polar_comm _ x.1 x.2, + smul_sub, LinearMap.smul_def, sub_add_sub_cancel, ← sub_eq_add_neg (Q x.1) (Q x.2)] + rw [← LinearMap.map_sub (⅟ 2 : Module.End R R), ← mul_sub, ← LinearMap.smul_def] + simp only [LinearMap.smul_def, half_moduleEnd_apply_eq_half_smul, smul_eq_mul, + invOf_mul_cancel_left'] /-! TODO: show that `QuadraticForm.toDualProd` is an `QuadraticForm.IsometryEquiv` From bcc3ed61725456f92fc9e6bd4fa23ef842176e77 Mon Sep 17 00:00:00 2001 From: yhtq <1414672068@qq.com> Date: Sat, 23 Nov 2024 08:16:22 +0000 Subject: [PATCH 227/829] feat: add Algebra.compHom and related lemma (#18404) Add some simple definition and lemmas to construct Algebra structure from Algebra on another equivalent ring. As preparation for the [following result](https://github.com/mbkybky/GaloisRamification/blob/1070c2f8c3359e891d29e30b3e7185d5abc78a86/GaloisRamification/ToMathlib/IsGalois.lean#L111): If `K` and `K'` are fraction rings of `A`, and `L` and `L'` are fraction rings of `B`, then `IsGalois K L` implies `IsGalois K' L'`. --- Mathlib/Algebra/Algebra/Defs.lean | 43 +++++++++++++++++++++++++++++++ 1 file changed, 43 insertions(+) diff --git a/Mathlib/Algebra/Algebra/Defs.lean b/Mathlib/Algebra/Algebra/Defs.lean index f21cc3fe2a3ef..40ac1bbeaf52e 100644 --- a/Mathlib/Algebra/Algebra/Defs.lean +++ b/Mathlib/Algebra/Algebra/Defs.lean @@ -175,6 +175,19 @@ def RingHom.toAlgebra' {R S} [CommSemiring R] [Semiring S] (i : R →+* S) smul_def' _ _ := rfl toRingHom := i +-- just simple lemmas for a declaration that is itself primed, no need for docstrings +set_option linter.docPrime false in +theorem RingHom.smul_toAlgebra' {R S} [CommSemiring R] [Semiring S] (i : R →+* S) + (h : ∀ c x, i c * x = x * i c) (r : R) (s : S) : + let _ := RingHom.toAlgebra' i h + r • s = i r * s := rfl + +set_option linter.docPrime false in +theorem RingHom.algebraMap_toAlgebra' {R S} [CommSemiring R] [Semiring S] (i : R →+* S) + (h : ∀ c x, i c * x = x * i c) : + @algebraMap R S _ _ (i.toAlgebra' h) = i := + rfl + /-- Creating an algebra from a morphism to a commutative semiring. -/ def RingHom.toAlgebra {R S} [CommSemiring R] [CommSemiring S] (i : R →+* S) : Algebra R S := i.toAlgebra' fun _ => mul_comm _ @@ -311,6 +324,36 @@ section end +section compHom + +variable (A) (f : S →+* R) + +/-- +Compose an `Algebra` with a `RingHom`, with action `f s • m`. + +This is the algebra version of `Module.compHom`. +-/ +abbrev compHom : Algebra S A where + smul s a := f s • a + toRingHom := (algebraMap R A).comp f + commutes' _ _ := Algebra.commutes _ _ + smul_def' _ _ := Algebra.smul_def _ _ + +theorem compHom_smul_def (s : S) (x : A) : + letI := compHom A f + s • x = f s • x := rfl + +theorem compHom_algebraMap_eq : + letI := compHom A f + algebraMap S A = (algebraMap R A).comp f := rfl + +theorem compHom_algebraMap_apply (s : S) : + letI := compHom A f + algebraMap S A s = (algebraMap R A) (f s) := rfl + +end compHom + + variable (R A) /-- The canonical ring homomorphism `algebraMap R A : R →+* A` for any `R`-algebra `A`, From 8eedc61a030954e3dad5a06bd750b29f54709126 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Sat, 23 Nov 2024 08:16:23 +0000 Subject: [PATCH 228/829] chore(discover-lean-pr-testing): better tracing, better output formatting (#19392) --- .github/workflows/discover-lean-pr-testing.yml | 13 ++++++++----- 1 file changed, 8 insertions(+), 5 deletions(-) diff --git a/.github/workflows/discover-lean-pr-testing.yml b/.github/workflows/discover-lean-pr-testing.yml index 2bee99c8e421d..8210a3e66daab 100644 --- a/.github/workflows/discover-lean-pr-testing.yml +++ b/.github/workflows/discover-lean-pr-testing.yml @@ -74,9 +74,12 @@ jobs: run: | # Use the PRs from the previous step PRS="${{ steps.get-prs.outputs.prs }}" + echo "$PRS" + echo "====================" echo "$PRS" | tr ' ' '\n' > prs.txt MATCHING_BRANCHES=$(git branch -r | grep -f prs.txt) echo "$MATCHING_BRANCHES" + echo "====================" # Initialize an empty variable to store branches with relevant diffs RELEVANT_BRANCHES="" @@ -89,25 +92,25 @@ jobs: # Check if the diff contains files other than the specified ones if echo "$DIFF_FILES" | grep -v -e 'lake-manifest.json' -e 'lakefile.lean' -e 'lean-toolchain'; then # If it does, add the branch to RELEVANT_BRANCHES - RELEVANT_BRANCHES="$RELEVANT_BRANCHES $BRANCH" + RELEVANT_BRANCHES="$RELEVANT_BRANCHES\n- $BRANCH" fi done # Output the relevant branches echo "'$RELEVANT_BRANCHES'" - echo "branches=$RELEVANT_BRANCHES" >> "$GITHUB_OUTPUT" + printf "branches<> "$GITHUB_OUTPUT" - name: Check if there are relevant branches id: check-branches run: | if [ -z "${{ steps.find-branches.outputs.branches }}" ]; then - echo "no_branches=true" >> "$GITHUB_ENV" + echo "branches_exist=false" >> "$GITHUB_ENV" else - echo "no_branches=false" >> "$GITHUB_ENV" + echo "branches_exist=true" >> "$GITHUB_ENV" fi - name: Send message on Zulip - if: env.no_branches == 'false' + if: env.branches_exist == 'true' uses: zulip/github-actions-zulip/send-message@v1 with: api-key: ${{ secrets.ZULIP_API_KEY }} From ba61bf3e0b903f7186a3861a6cf1d704a1bed7c0 Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Sat, 23 Nov 2024 08:46:42 +0000 Subject: [PATCH 229/829] chore(CategoryTheory.ConcreteCategory): inline `ConcreteCategory.forget` (#19217) We already make `forget` `reducible` and further declarations downstream of it `abbrev`'s. It makes sense to `inline` it also to avoid Lean noticing it. --- Mathlib/CategoryTheory/ConcreteCategory/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/ConcreteCategory/Basic.lean b/Mathlib/CategoryTheory/ConcreteCategory/Basic.lean index 38882d4e9f448..ebeea719fe2dc 100644 --- a/Mathlib/CategoryTheory/ConcreteCategory/Basic.lean +++ b/Mathlib/CategoryTheory/ConcreteCategory/Basic.lean @@ -55,7 +55,7 @@ class ConcreteCategory (C : Type u) [Category.{v} C] where /-- That functor is faithful -/ [forget_faithful : forget.Faithful] -attribute [reducible] ConcreteCategory.forget +attribute [inline, reducible] ConcreteCategory.forget attribute [instance] ConcreteCategory.forget_faithful /-- The forgetful functor from a concrete category to `Type u`. -/ From 4fd77b8b24b94a4d2eb8e4409675e0542f8a415e Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 23 Nov 2024 09:01:50 +0000 Subject: [PATCH 230/829] chore: missing MvPolynomial.eval lemmas (#19356) Revealed by a new user question on [zulip](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/Adding.20more.20examples.20to.20mathlib/near/483822621). Co-authored-by: Eric Wieser --- Mathlib/Algebra/MvPolynomial/Basic.lean | 29 +++++++++++++++++++++++++ 1 file changed, 29 insertions(+) diff --git a/Mathlib/Algebra/MvPolynomial/Basic.lean b/Mathlib/Algebra/MvPolynomial/Basic.lean index b73915560d966..b649bbdca1c39 100644 --- a/Mathlib/Algebra/MvPolynomial/Basic.lean +++ b/Mathlib/Algebra/MvPolynomial/Basic.lean @@ -919,6 +919,14 @@ theorem eval₂_C (a) : (C a).eval₂ f g = f a := by theorem eval₂_one : (1 : MvPolynomial σ R).eval₂ f g = 1 := (eval₂_C _ _ _).trans f.map_one +@[simp] theorem eval₂_natCast (n : Nat) : (n : MvPolynomial σ R).eval₂ f g = n := + (eval₂_C _ _ _).trans (map_natCast f n) + +-- See note [no_index around OfNat.ofNat] +@[simp] theorem eval₂_ofNat (n : Nat) [n.AtLeastTwo] : + (no_index (OfNat.ofNat n) : MvPolynomial σ R).eval₂ f g = OfNat.ofNat n := + eval₂_natCast f g n + @[simp] theorem eval₂_X (n) : (X n).eval₂ f g = g n := by simp [eval₂_monomial, f.map_one, X, prod_single_index, pow_one] @@ -1069,6 +1077,11 @@ theorem eval_C : ∀ a, eval f (C a) = a := theorem eval_X : ∀ n, eval f (X n) = f n := eval₂_X _ _ +-- See note [no_index around OfNat.ofNat] +@[simp] theorem eval_ofNat (n : Nat) [n.AtLeastTwo] : + (no_index (OfNat.ofNat n) : MvPolynomial σ R).eval f = OfNat.ofNat n := + map_ofNat _ n + @[simp] theorem smul_eval (x) (p : MvPolynomial σ R) (s) : eval x (s • p) = s * eval x p := by rw [smul_eq_C_mul, (eval x).map_mul, eval_C] @@ -1129,6 +1142,11 @@ theorem map_monomial (s : σ →₀ ℕ) (a : R) : map f (monomial s a) = monomi theorem map_C : ∀ a : R, map f (C a : MvPolynomial σ R) = C (f a) := map_monomial _ _ +-- See note [no_index around OfNat.ofNat] +@[simp] protected theorem map_ofNat (n : Nat) [n.AtLeastTwo] : + (no_index (OfNat.ofNat n) : MvPolynomial σ R).map f = OfNat.ofNat n := + _root_.map_ofNat _ _ + @[simp] theorem map_X : ∀ n : σ, map f (X n : MvPolynomial σ R) = X n := eval₂_X _ _ @@ -1344,6 +1362,11 @@ theorem aeval_X (s : σ) : aeval f (X s : MvPolynomial _ R) = f s := theorem aeval_C (r : R) : aeval f (C r) = algebraMap R S₁ r := eval₂_C _ _ _ +-- See note [no_index around OfNat.ofNat] +@[simp] theorem aeval_ofNat (n : Nat) [n.AtLeastTwo] : + aeval f (no_index (OfNat.ofNat n) : MvPolynomial σ R) = OfNat.ofNat n := + map_ofNat _ _ + theorem aeval_unique (φ : MvPolynomial σ R →ₐ[R] S₁) : φ = aeval (φ ∘ X) := by ext i simp @@ -1475,6 +1498,12 @@ theorem aevalTower_X (i : σ) : aevalTower g y (X i) = y i := theorem aevalTower_C (x : R) : aevalTower g y (C x) = g x := eval₂_C _ _ _ +-- See note [no_index around OfNat.ofNat] +@[simp] +theorem aevalTower_ofNat (n : Nat) [n.AtLeastTwo] : + aevalTower g y (no_index (OfNat.ofNat n) : MvPolynomial σ R) = OfNat.ofNat n := + _root_.map_ofNat _ _ + @[simp] theorem aevalTower_comp_C : (aevalTower g y : MvPolynomial σ R →+* A).comp C = g := RingHom.ext <| aevalTower_C _ _ From 25a3b654353e3d52ffd0c255f313600fcea9e34a Mon Sep 17 00:00:00 2001 From: Christopher Hoskin Date: Sat, 23 Nov 2024 09:23:50 +0000 Subject: [PATCH 231/829] feature(Order/SupClosed): Add lemmas for inserting upper or lower bounds into a Sup/Inf closed set (#18740) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add lemmas for inserting upper or lower bounds into a Sup/Inf closed set. e.g. When `s` is sup closed, so is `s` with `⊤` or `⊥` appended. Similarly for `s` inf closed. Co-authored-by: Christopher Hoskin --- Mathlib/Order/SupClosed.lean | 22 ++++++++++++++++++++++ 1 file changed, 22 insertions(+) diff --git a/Mathlib/Order/SupClosed.lean b/Mathlib/Order/SupClosed.lean index 61ddd8850360c..84fdcccc76f9d 100644 --- a/Mathlib/Order/SupClosed.lean +++ b/Mathlib/Order/SupClosed.lean @@ -79,6 +79,17 @@ lemma supClosed_pi {ι : Type*} {α : ι → Type*} [∀ i, SemilatticeSup (α i {t : ∀ i, Set (α i)} (ht : ∀ i ∈ s, SupClosed (t i)) : SupClosed (s.pi t) := fun _a ha _b hb _i hi ↦ ht _ hi (ha _ hi) (hb _ hi) +lemma SupClosed.insert_upperBounds {s : Set α} {a : α} (hs : SupClosed s) (ha : a ∈ upperBounds s) : + SupClosed (insert a s) := by + rw [SupClosed] + aesop + +lemma SupClosed.insert_lowerBounds {s : Set α} {a : α} (h : SupClosed s) (ha : a ∈ lowerBounds s) : + SupClosed (insert a s) := by + rw [SupClosed] + have ha' : ∀ b ∈ s, a ≤ b := fun _ a ↦ ha a + aesop + end Set section Finset @@ -144,6 +155,17 @@ lemma infClosed_pi {ι : Type*} {α : ι → Type*} [∀ i, SemilatticeInf (α i {t : ∀ i, Set (α i)} (ht : ∀ i ∈ s, InfClosed (t i)) : InfClosed (s.pi t) := fun _a ha _b hb _i hi ↦ ht _ hi (ha _ hi) (hb _ hi) +lemma InfClosed.insert_upperBounds {s : Set α} {a : α} (hs : InfClosed s) (ha : a ∈ upperBounds s) : + InfClosed (insert a s) := by + rw [InfClosed] + have ha' : ∀ b ∈ s, b ≤ a := fun _ a ↦ ha a + aesop + +lemma InfClosed.insert_lowerBounds {s : Set α} {a : α} (h : InfClosed s) (ha : a ∈ lowerBounds s) : + InfClosed (insert a s) := by + rw [InfClosed] + aesop + end Set section Finset From d1d52a625399fe9f91ef60a29818f4005053e70c Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Sat, 23 Nov 2024 09:53:38 +0000 Subject: [PATCH 232/829] chore: update Mathlib dependencies 2024-11-23 (#19396) 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 762bb1280e3ba..6e18dc8970e13 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "8b52587ff32e2e443cce6109b5305341289339e7", + "rev": "0dc51ac7947ff6aa2c16bcffb64c46c7149d1276", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From 0b15a0dc917eb5bdc9450c801c99423944d6045e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 23 Nov 2024 11:01:40 +0000 Subject: [PATCH 233/829] feat: independence of functions under conditioning (#17915) From PFR --- Mathlib/Analysis/SpecificLimits/Basic.lean | 2 +- Mathlib/Data/ENNReal/Inv.lean | 10 ++++ Mathlib/MeasureTheory/Function/LpSpace.lean | 2 +- Mathlib/MeasureTheory/Integral/Bochner.lean | 2 +- .../Probability/ConditionalProbability.lean | 4 +- Mathlib/Probability/Independence/Basic.lean | 57 +++++++++++++++++++ Mathlib/Probability/Independence/Kernel.lean | 53 +++++++++++++++++ 7 files changed, 125 insertions(+), 5 deletions(-) diff --git a/Mathlib/Analysis/SpecificLimits/Basic.lean b/Mathlib/Analysis/SpecificLimits/Basic.lean index 3d759d9f3fb33..e9bbbad7a0c52 100644 --- a/Mathlib/Analysis/SpecificLimits/Basic.lean +++ b/Mathlib/Analysis/SpecificLimits/Basic.lean @@ -623,7 +623,7 @@ theorem tendsto_factorial_div_pow_self_atTop : refine (eventually_gt_atTop 0).mono fun n hn ↦ ?_ rcases Nat.exists_eq_succ_of_ne_zero hn.ne.symm with ⟨k, rfl⟩ rw [← prod_range_add_one_eq_factorial, pow_eq_prod_const, div_eq_mul_inv, ← inv_eq_one_div, - prod_natCast, Nat.cast_succ, ← prod_inv_distrib, ← prod_mul_distrib, + prod_natCast, Nat.cast_succ, ← Finset.prod_inv_distrib, ← prod_mul_distrib, Finset.prod_range_succ'] simp only [prod_range_succ', one_mul, Nat.cast_add, zero_add, Nat.cast_one] refine diff --git a/Mathlib/Data/ENNReal/Inv.lean b/Mathlib/Data/ENNReal/Inv.lean index 36003f724d580..bccdde7f1f4ab 100644 --- a/Mathlib/Data/ENNReal/Inv.lean +++ b/Mathlib/Data/ENNReal/Inv.lean @@ -173,6 +173,16 @@ protected theorem inv_div {a b : ℝ≥0∞} (htop : b ≠ ∞ ∨ a ≠ ∞) (h rw [← ENNReal.inv_ne_top] at hzero rw [ENNReal.div_eq_inv_mul, ENNReal.div_eq_inv_mul, ENNReal.mul_inv htop hzero, mul_comm, inv_inv] +lemma prod_inv_distrib {ι : Type*} {f : ι → ℝ≥0∞} {s : Finset ι} + (hf : s.toSet.Pairwise fun i j ↦ f i ≠ 0 ∨ f j ≠ ∞) : (∏ i ∈ s, f i)⁻¹ = ∏ i ∈ s, (f i)⁻¹ := by + induction' s using Finset.cons_induction with i s hi ih + · simp + simp [← ih (hf.mono <| by simp)] + refine ENNReal.mul_inv (not_or_of_imp fun hi₀ ↦ prod_ne_top fun j hj ↦ ?_) + (not_or_of_imp fun hi₀ ↦ Finset.prod_ne_zero_iff.2 fun j hj ↦ ?_) + · exact imp_iff_not_or.2 (hf (by simp) (by simp [hj]) <| .symm <| ne_of_mem_of_not_mem hj hi) hi₀ + · exact imp_iff_not_or.2 (hf (by simp [hj]) (by simp) <| ne_of_mem_of_not_mem hj hi).symm hi₀ + protected theorem mul_div_mul_left (a b : ℝ≥0∞) (hc : c ≠ 0) (hc' : c ≠ ⊤) : c * a / (c * b) = a / b := by rw [div_eq_mul_inv, div_eq_mul_inv, ENNReal.mul_inv (Or.inl hc) (Or.inl hc'), mul_mul_mul_comm, diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index 5b9c9290dac61..582ddfc879f14 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -1192,7 +1192,7 @@ def compLpₗ (L : E →L[𝕜] F) : Lp E p μ →ₗ[𝕜] Lp F p μ where ext1 filter_upwards [Lp.coeFn_smul c f, coeFn_compLp L (c • f), Lp.coeFn_smul c (L.compLp f), coeFn_compLp L f] with _ ha1 ha2 ha3 ha4 - simp only [ha1, ha2, ha3, ha4, map_smul, Pi.smul_apply] + simp only [ha1, ha2, ha3, ha4, _root_.map_smul, Pi.smul_apply] /-- Composing `f : Lp E p μ` with `L : E →L[𝕜] F`, seen as a continuous `𝕜`-linear map on `Lp E p μ`. See also the similar diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index c85ae9df03c29..2cf72496737ad 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -638,7 +638,7 @@ theorem integral_sub (f g : α →₁[μ] E) : integral (f - g) = integral f - i theorem integral_smul (c : 𝕜) (f : α →₁[μ] E) : integral (c • f) = c • integral f := by simp only [integral] show (integralCLM' (E := E) 𝕜) (c • f) = c • (integralCLM' (E := E) 𝕜) f - exact map_smul (integralCLM' (E := E) 𝕜) c f + exact _root_.map_smul (integralCLM' (E := E) 𝕜) c f local notation "Integral" => @integralCLM α E _ _ μ _ _ diff --git a/Mathlib/Probability/ConditionalProbability.lean b/Mathlib/Probability/ConditionalProbability.lean index 6dcf1ba7c951b..4e4d89fc5ac91 100644 --- a/Mathlib/Probability/ConditionalProbability.lean +++ b/Mathlib/Probability/ConditionalProbability.lean @@ -70,8 +70,8 @@ and scaled by the inverse of `μ s` (to make it a probability measure): def cond (s : Set Ω) : Measure Ω := (μ s)⁻¹ • μ.restrict s -@[inherit_doc] scoped notation:max μ "[" t " | " s "]" => ProbabilityTheory.cond μ s t @[inherit_doc] scoped notation:max μ "[|" s "]" => ProbabilityTheory.cond μ s +@[inherit_doc cond] scoped notation3:max μ "[" t " | " s "]" => ProbabilityTheory.cond μ s t /-- The conditional probability measure of measure `μ` on `{ω | X ω ∈ s}`. @@ -172,7 +172,7 @@ theorem inter_pos_of_cond_ne_zero (hms : MeasurableSet s) (hcst : μ[t|s] ≠ 0) simp [hms, Set.inter_comm, cond] lemma cond_pos_of_inter_ne_zero [IsFiniteMeasure μ] (hms : MeasurableSet s) (hci : μ (s ∩ t) ≠ 0) : - 0 < μ[|s] t := by + 0 < μ[t | s] := by rw [cond_apply hms] refine ENNReal.mul_pos ?_ hci exact ENNReal.inv_ne_zero.mpr (measure_ne_top _ _) diff --git a/Mathlib/Probability/Independence/Basic.lean b/Mathlib/Probability/Independence/Basic.lean index cfa5d0c19e7b7..a7ef0fcd188d8 100644 --- a/Mathlib/Probability/Independence/Basic.lean +++ b/Mathlib/Probability/Independence/Basic.lean @@ -721,4 +721,61 @@ theorem iIndepSet.iIndepFun_indicator [Zero β] [One β] {m : MeasurableSpace β end IndepFun +variable {ι Ω α β : Type*} {mΩ : MeasurableSpace Ω} {mα : MeasurableSpace α} + {mβ : MeasurableSpace β} {μ : Measure Ω} {X : ι → Ω → α} {Y : ι → Ω → β} {f : _ → Set Ω} + {t : ι → Set β} {s : Finset ι} + +/-- The probability of an intersection of preimages conditioning on another intersection factors +into a product. -/ +lemma cond_iInter [Finite ι] (hY : ∀ i, Measurable (Y i)) + (hindep : iIndepFun (fun _ ↦ mα.prod mβ) (fun i ω ↦ (X i ω, Y i ω)) μ) + (hf : ∀ i ∈ s, MeasurableSet[mα.comap (X i)] (f i)) + (hy : ∀ i ∉ s, μ (Y i ⁻¹' t i) ≠ 0) (ht : ∀ i, MeasurableSet (t i)) : + μ[⋂ i ∈ s, f i | ⋂ i, Y i ⁻¹' t i] = ∏ i ∈ s, μ[f i | Y i in t i] := by + have : IsProbabilityMeasure (μ : Measure Ω) := hindep.isProbabilityMeasure + classical + cases nonempty_fintype ι + let g (i' : ι) := if i' ∈ s then Y i' ⁻¹' t i' ∩ f i' else Y i' ⁻¹' t i' + calc + _ = (μ (⋂ i, Y i ⁻¹' t i))⁻¹ * μ ((⋂ i, Y i ⁻¹' t i) ∩ ⋂ i ∈ s, f i) := by + rw [cond_apply]; exact .iInter fun i ↦ hY i (ht i) + _ = (μ (⋂ i, Y i ⁻¹' t i))⁻¹ * μ (⋂ i, g i) := by + congr + calc + _ = (⋂ i, Y i ⁻¹' t i) ∩ ⋂ i, if i ∈ s then f i else .univ := by + congr + simp only [Set.iInter_ite, Set.iInter_univ, Set.inter_univ] + _ = ⋂ i, Y i ⁻¹' t i ∩ (if i ∈ s then f i else .univ) := by rw [Set.iInter_inter_distrib] + _ = _ := Set.iInter_congr fun i ↦ by by_cases hi : i ∈ s <;> simp [hi, g] + _ = (∏ i, μ (Y i ⁻¹' t i))⁻¹ * μ (⋂ i, g i) := by + rw [hindep.meas_iInter] + exact fun i ↦ ⟨.univ ×ˢ t i, MeasurableSet.univ.prod (ht _), by ext; simp [eq_comm]⟩ + _ = (∏ i, μ (Y i ⁻¹' t i))⁻¹ * ∏ i, μ (g i) := by + rw [hindep.meas_iInter] + intro i + by_cases hi : i ∈ s <;> simp only [hi, ↓reduceIte, g] + · obtain ⟨A, hA, hA'⟩ := hf i hi + exact .inter ⟨.univ ×ˢ t i, MeasurableSet.univ.prod (ht _), by ext; simp [eq_comm]⟩ + ⟨A ×ˢ Set.univ, hA.prod .univ, by ext; simp [← hA']⟩ + · exact ⟨.univ ×ˢ t i, MeasurableSet.univ.prod (ht _), by ext; simp [eq_comm]⟩ + _ = ∏ i, (μ (Y i ⁻¹' t i))⁻¹ * μ (g i) := by + rw [Finset.prod_mul_distrib, ENNReal.prod_inv_distrib] + exact fun _ _ i _ _ ↦ .inr <| measure_ne_top _ _ + _ = ∏ i, if i ∈ s then μ[f i | Y i ⁻¹' t i] else 1 := by + refine Finset.prod_congr rfl fun i _ ↦ ?_ + by_cases hi : i ∈ s + · simp only [hi, ↓reduceIte, g, cond_apply (hY i (ht i))] + · simp only [hi, ↓reduceIte, g, ENNReal.inv_mul_cancel (hy i hi) (measure_ne_top μ _)] + _ = _ := by simp + +lemma iIndepFun.cond [Finite ι] (hY : ∀ i, Measurable (Y i)) + (hindep : iIndepFun (fun _ ↦ mα.prod mβ) (fun i ω ↦ (X i ω, Y i ω)) μ) + (hy : ∀ i, μ (Y i ⁻¹' t i) ≠ 0) (ht : ∀ i, MeasurableSet (t i)) : + iIndepFun (fun _ ↦ mα) X μ[|⋂ i, Y i ⁻¹' t i] := by + rw [iIndepFun_iff] + intro s f hf + convert cond_iInter hY hindep hf (fun i _ ↦ hy _) ht using 2 with i hi + simpa using cond_iInter hY hindep (fun j hj ↦ hf _ <| Finset.mem_singleton.1 hj ▸ hi) + (fun i _ ↦ hy _) ht + end ProbabilityTheory diff --git a/Mathlib/Probability/Independence/Kernel.lean b/Mathlib/Probability/Independence/Kernel.lean index 85ed49143dcc1..3f4c1971fcafe 100644 --- a/Mathlib/Probability/Independence/Kernel.lean +++ b/Mathlib/Probability/Independence/Kernel.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ import Mathlib.MeasureTheory.Constructions.Pi +import Mathlib.Probability.ConditionalProbability import Mathlib.Probability.Kernel.Basic import Mathlib.Tactic.Peel @@ -1212,4 +1213,56 @@ theorem iIndepSet.iIndepFun_indicator [Zero β] [One β] {m : MeasurableSpace β end IndepFun +variable {ι Ω α β : Type*} {mΩ : MeasurableSpace Ω} {mα : MeasurableSpace α} + {mβ : MeasurableSpace β} {κ : Kernel α Ω} {μ : Measure α} {X : ι → Ω → α} {Y : ι → Ω → β} + {f : _ → Set Ω} {t : ι → Set β} {s : Finset ι} + +/-- The probability of an intersection of preimages conditioning on another intersection factors +into a product. -/ +lemma iIndepFun.cond_iInter [Finite ι] (hY : ∀ i, Measurable (Y i)) + (hindep : iIndepFun (fun _ ↦ mα.prod mβ) (fun i ω ↦ (X i ω, Y i ω)) κ μ) + (hf : ∀ i ∈ s, MeasurableSet[mα.comap (X i)] (f i)) + (hy : ∀ᵐ a ∂μ, ∀ i ∉ s, κ a (Y i ⁻¹' t i) ≠ 0) (ht : ∀ i, MeasurableSet (t i)) : + ∀ᵐ a ∂μ, (κ a)[⋂ i ∈ s, f i | ⋂ i, Y i ⁻¹' t i] = ∏ i ∈ s, (κ a)[f i | Y i in t i] := by + classical + cases nonempty_fintype ι + let g (i' : ι) := if i' ∈ s then Y i' ⁻¹' t i' ∩ f i' else Y i' ⁻¹' t i' + have hYt i : MeasurableSet[(mα.prod mβ).comap fun ω ↦ (X i ω, Y i ω)] (Y i ⁻¹' t i) := + ⟨.univ ×ˢ t i, .prod .univ (ht _), by ext; simp [eq_comm]⟩ + have hg i : MeasurableSet[(mα.prod mβ).comap fun ω ↦ (X i ω, Y i ω)] (g i) := by + by_cases hi : i ∈ s <;> simp only [hi, ↓reduceIte, g] + · obtain ⟨A, hA, hA'⟩ := hf i hi + exact (hYt _).inter ⟨A ×ˢ .univ, hA.prod .univ, by ext; simp [← hA']⟩ + · exact hYt _ + filter_upwards [hy, hindep.ae_isProbabilityMeasure, hindep.meas_iInter hYt, hindep.meas_iInter hg] + with a hy _ hYt hg + calc + _ = (κ a (⋂ i, Y i ⁻¹' t i))⁻¹ * κ a ((⋂ i, Y i ⁻¹' t i) ∩ ⋂ i ∈ s, f i) := by + rw [cond_apply]; exact .iInter fun i ↦ hY i (ht i) + _ = (κ a (⋂ i, Y i ⁻¹' t i))⁻¹ * κ a (⋂ i, g i) := by + congr + calc + _ = (⋂ i, Y i ⁻¹' t i) ∩ ⋂ i, if i ∈ s then f i else .univ := by + congr + simp only [Set.iInter_ite, Set.iInter_univ, Set.inter_univ] + _ = ⋂ i, Y i ⁻¹' t i ∩ (if i ∈ s then f i else .univ) := by rw [Set.iInter_inter_distrib] + _ = _ := Set.iInter_congr fun i ↦ by by_cases hi : i ∈ s <;> simp [hi, g] + _ = (∏ i, κ a (Y i ⁻¹' t i))⁻¹ * κ a (⋂ i, g i) := by + rw [hYt] + _ = (∏ i, κ a (Y i ⁻¹' t i))⁻¹ * ∏ i, κ a (g i) := by + rw [hg] + _ = ∏ i, (κ a (Y i ⁻¹' t i))⁻¹ * κ a (g i) := by + rw [Finset.prod_mul_distrib, ENNReal.prod_inv_distrib] + exact fun _ _ i _ _ ↦ .inr <| measure_ne_top _ _ + _ = ∏ i, if i ∈ s then (κ a)[f i | Y i ⁻¹' t i] else 1 := by + refine Finset.prod_congr rfl fun i _ ↦ ?_ + by_cases hi : i ∈ s + · simp only [hi, ↓reduceIte, g, cond_apply (hY i (ht i))] + · simp only [hi, ↓reduceIte, g, ENNReal.inv_mul_cancel (hy i hi) (measure_ne_top _ _)] + _ = _ := by simp + +-- TODO: We can't state `Kernel.iIndepFun.cond` (the `Kernel` analogue of +-- `ProbabilityTheory.iIndepFun.cond`) because we don't have a version of `ProbabilityTheory.cond` +-- for kernels + end ProbabilityTheory.Kernel From 09a053e649888befa819426d00878d7f3aaca830 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Sat, 23 Nov 2024 13:04:00 +0000 Subject: [PATCH 234/829] feat(NumberTheory/LSeries/PrimesInAP): von Mangoldt restricted to a residue class (#19368) This is the next step on the way to **Dirichlet's Theorem**. It defines `ArithmeticFunction.vonMangoldt.residueClass` as the function that is the von Mangoldt function on a given residue class and zero outside and proves some facts about it (in particular that it is a linear combination of twists of the von Mangoldt function by Dirichlet characters). See [here](https://leanprover.zulipchat.com/#narrow/channel/144837-PR-reviews/topic/Prerequisites.20for.20PNT.20and.20Dirichlet's.20Thm/near/483785853) on Zulip. The large import increase is caused by the need for more material to be able to state and prove the new things. --- Mathlib/NumberTheory/LSeries/PrimesInAP.lean | 94 ++++++++++++++++++-- 1 file changed, 89 insertions(+), 5 deletions(-) diff --git a/Mathlib/NumberTheory/LSeries/PrimesInAP.lean b/Mathlib/NumberTheory/LSeries/PrimesInAP.lean index 0c1fc1a1ebf26..0740c4659750a 100644 --- a/Mathlib/NumberTheory/LSeries/PrimesInAP.lean +++ b/Mathlib/NumberTheory/LSeries/PrimesInAP.lean @@ -3,9 +3,10 @@ Copyright (c) 2024 Michael Stoll. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Stoll -/ -import Mathlib.Data.Nat.Factorization.PrimePow -import Mathlib.Topology.Algebra.InfiniteSum.Constructions -import Mathlib.Topology.Algebra.InfiniteSum.NatInt +import Mathlib.NumberTheory.DirichletCharacter.Orthogonality +import Mathlib.NumberTheory.LSeries.DirichletContinuation +import Mathlib.NumberTheory.LSeries.Linearity +import Mathlib.RingTheory.RootsOfUnity.AlgebraicallyClosed /-! # Dirichlet's Theorem on primes in arithmetic progression @@ -17,8 +18,8 @@ and `a : ZMod q` is invertible, then there are infinitely many prime numbers `p` This will be done in the following steps: 1. Some auxiliary lemmas on infinite products and sums -2. (TODO) Results on the von Mangoldt function restricted to a residue class -3. (TODO) Results on its L-series +2. Results on the von Mangoldt function restricted to a residue class +3. (TODO) Results on its L-series and an auxiliary function related to it 4. (TODO) Derivation of Dirichlet's Theorem -/ @@ -65,3 +66,86 @@ lemma tprod_eq_tprod_primes_mul_tprod_primes_of_mulSupport_subset_prime_powers { Function.Injective.prodMap (fun ⦃_ _⦄ a ↦ a) <| add_left_injective 1 end auxiliary + +/-! +### The L-series of the von Mangoldt function restricted to a residue class +-/ + +section arith_prog + +namespace ArithmeticFunction.vonMangoldt + +open Complex LSeries DirichletCharacter + +open scoped LSeries.notation + +variable {q : ℕ} (a : ZMod q) + +/-- The von Mangoldt function restricted to the residue class `a` mod `q`. -/ +noncomputable abbrev residueClass : ℕ → ℝ := + {n : ℕ | (n : ZMod q) = a}.indicator (vonMangoldt ·) + +lemma residueClass_nonneg (n : ℕ) : 0 ≤ residueClass a n := + Set.indicator_apply_nonneg fun _ ↦ vonMangoldt_nonneg + +lemma residueClass_le (n : ℕ) : residueClass a n ≤ vonMangoldt n := + Set.indicator_apply_le' (fun _ ↦ le_rfl) (fun _ ↦ vonMangoldt_nonneg) + +@[simp] +lemma residueClass_apply_zero : residueClass a 0 = 0 := by + simp only [Set.indicator_apply_eq_zero, Set.mem_setOf_eq, Nat.cast_zero, map_zero, ofReal_zero, + implies_true] + +lemma abscissaOfAbsConv_residueClass_le_one : + abscissaOfAbsConv ↗(residueClass a) ≤ 1 := by + refine abscissaOfAbsConv_le_of_forall_lt_LSeriesSummable fun y hy ↦ ?_ + unfold LSeriesSummable + have := LSeriesSummable_vonMangoldt <| show 1 < (y : ℂ).re by simp only [ofReal_re, hy] + convert this.indicator {n : ℕ | (n : ZMod q) = a} + ext1 n + by_cases hn : (n : ZMod q) = a + · simp +contextual only [term, Set.indicator, Set.mem_setOf_eq, hn, ↓reduceIte, apply_ite, + ite_self] + · simp +contextual only [term, Set.mem_setOf_eq, hn, not_false_eq_true, Set.indicator_of_not_mem, + ofReal_zero, zero_div, ite_self] + +variable [NeZero q] {a} + +/-- We can express `ArithmeticFunction.vonMangoldt.residueClass` as a linear combination +of twists of the von Mangoldt function with Dirichlet charaters. -/ +lemma residueClass_apply (ha : IsUnit a) (n : ℕ) : + residueClass a n = + (q.totient : ℂ)⁻¹ * ∑ χ : DirichletCharacter ℂ q, χ a⁻¹ * χ n * vonMangoldt n := by + rw [eq_inv_mul_iff_mul_eq₀ <| mod_cast (Nat.totient_pos.mpr q.pos_of_neZero).ne'] + simp +contextual only [residueClass, Set.indicator_apply, Set.mem_setOf_eq, apply_ite, + ofReal_zero, mul_zero, ← Finset.sum_mul, sum_char_inv_mul_char_eq ℂ ha n, eq_comm (a := a), + ite_mul, zero_mul, ↓reduceIte, ite_self] + +/-- We can express `ArithmeticFunction.vonMangoldt.residueClass` as a linear combination +of twists of the von Mangoldt function with Dirichlet charaters. -/ +lemma residueClass_eq (ha : IsUnit a) : + ↗(residueClass a) = (q.totient : ℂ)⁻¹ • + ∑ χ : DirichletCharacter ℂ q, χ a⁻¹ • (fun n : ℕ ↦ χ n * vonMangoldt n) := by + ext1 n + simpa only [Pi.smul_apply, Finset.sum_apply, smul_eq_mul, ← mul_assoc] + using residueClass_apply ha n + +/-- The L-series of the von Mangoldt function restricted to the residue class `a` mod `q` +with `a` invertible in `ZMod q` is a linear combination of logarithmic derivatives of +L-functions of the Dirichlet characters mod `q` (on `re s > 1`). -/ +lemma LSeries_residueClass_eq (ha : IsUnit a) {s : ℂ} (hs : 1 < s.re) : + LSeries ↗(residueClass a) s = + -(q.totient : ℂ)⁻¹ * ∑ χ : DirichletCharacter ℂ q, χ a⁻¹ * + (deriv (LFunction χ) s / LFunction χ s) := by + simp only [deriv_LFunction_eq_deriv_LSeries _ hs, LFunction_eq_LSeries _ hs, neg_mul, ← mul_neg, + ← Finset.sum_neg_distrib, ← neg_div, ← LSeries_twist_vonMangoldt_eq _ hs] + rw [eq_inv_mul_iff_mul_eq₀ <| mod_cast (Nat.totient_pos.mpr q.pos_of_neZero).ne'] + simp_rw [← LSeries_smul, + ← LSeries_sum <| fun χ _ ↦ (LSeriesSummable_twist_vonMangoldt χ hs).smul _] + refine LSeries_congr s fun {n} _ ↦ ?_ + simp only [Pi.smul_apply, residueClass_apply ha, smul_eq_mul, ← mul_assoc, + mul_inv_cancel_of_invertible, one_mul, Finset.sum_apply, Pi.mul_apply] + +end ArithmeticFunction.vonMangoldt + +end arith_prog From 64cf70aebf299bdf23a33ae80e1f1025193e0062 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 23 Nov 2024 13:12:26 +0000 Subject: [PATCH 235/829] feat: `DecidableEq` instance for quotient groups (#19258) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Hard to believe this is missing, but you can check for yourself: ``` import Mathlib variable {α : Type*} [Group α] (s : Subgroup α) [DecidablePred (· ∈ s)] #synth DecidableEq (α ⧸ s) ``` From GrowthInGroups --- Mathlib/GroupTheory/Coset/Defs.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/Mathlib/GroupTheory/Coset/Defs.lean b/Mathlib/GroupTheory/Coset/Defs.lean index f43753df9ac06..bd49ccf9a7568 100644 --- a/Mathlib/GroupTheory/Coset/Defs.lean +++ b/Mathlib/GroupTheory/Coset/Defs.lean @@ -89,6 +89,10 @@ instance leftRelDecidable [DecidablePred (· ∈ s)] : DecidableRel (leftRel s). instance instHasQuotientSubgroup : HasQuotient α (Subgroup α) := ⟨fun s => Quotient (leftRel s)⟩ +@[to_additive] +instance [DecidablePred (· ∈ s)] : DecidableEq (α ⧸ s) := + @Quotient.decidableEq _ _ (leftRelDecidable _) + /-- The equivalence relation corresponding to the partition of a group by right cosets of a subgroup. -/ @[to_additive "The equivalence relation corresponding to the partition of a group by right cosets From 142886c5f379dacb026bdc0fd88a4b6dbee48bec Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Sat, 23 Nov 2024 13:12:27 +0000 Subject: [PATCH 236/829] chore(RingTheory/PrincipalIdealDomain): update module docs (#19361) The module documentation for `RingTheory.PrincipalIdealDomain` was out of date, listing definitions that have been renamed and moved. --- Mathlib/RingTheory/PrincipalIdealDomain.lean | 7 ++++--- 1 file changed, 4 insertions(+), 3 deletions(-) diff --git a/Mathlib/RingTheory/PrincipalIdealDomain.lean b/Mathlib/RingTheory/PrincipalIdealDomain.lean index 7cde766527865..3a675457d9104 100644 --- a/Mathlib/RingTheory/PrincipalIdealDomain.lean +++ b/Mathlib/RingTheory/PrincipalIdealDomain.lean @@ -14,20 +14,21 @@ import Mathlib.RingTheory.Noetherian.UniqueFactorizationDomain A principal ideal ring (PIR) is a ring in which all left ideals are principal. A principal ideal domain (PID) is an integral domain which is a principal ideal ring. +The definition of `IsPrincipalIdealRing` can be found in `Mathlib.RingTheory.Ideal.Span`. + # Main definitions Note that for principal ideal domains, one should use `[IsDomain R] [IsPrincipalIdealRing R]`. There is no explicit definition of a PID. -Theorems about PID's are in the `principal_ideal_ring` namespace. +Theorems about PID's are in the `PrincipalIdealRing` namespace. -- `IsPrincipalIdealRing`: a predicate on rings, saying that every left ideal is principal. - `IsBezout`: the predicate saying that every finitely generated left ideal is principal. - `generator`: a generator of a principal ideal (or more generally submodule) - `to_uniqueFactorizationMonoid`: a PID is a unique factorization domain # Main results -- `to_maximal_ideal`: a non-zero prime ideal in a PID is maximal. +- `Ideal.IsPrime.to_maximal_ideal`: a non-zero prime ideal in a PID is maximal. - `EuclideanDomain.to_principal_ideal_domain` : a Euclidean domain is a PID. - `IsBezout.nonemptyGCDMonoid`: Every Bézout domain is a GCD domain. From 22a386f051ff3747f95dc507f0b80ed394ec9b2d Mon Sep 17 00:00:00 2001 From: Ahmad Alkhalawi Date: Sat, 23 Nov 2024 13:41:09 +0000 Subject: [PATCH 237/829] feat(RingTheory/DedekindDomain/IntegralClosure): Added `IsIntegralClosure.finite` (#18842) Let A be a Noetherian integrally closed domain with fraction field K. Let L/K be a finite separable field extension. Then the integral closure of A in L is finite over A. The proof immediately follows from `IsIntegralClosure.isNoetherian`. Co-authored-by: Ahmad Alkhalawi <46321632+4hma4d@users.noreply.github.com> --- Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean b/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean index 77679db0fe1c6..f38e19b607e41 100644 --- a/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean +++ b/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean @@ -185,6 +185,14 @@ theorem IsIntegralClosure.isNoetherianRing [IsIntegrallyClosed A] [IsNoetherianR IsNoetherianRing C := isNoetherianRing_iff.mpr <| isNoetherian_of_tower A (IsIntegralClosure.isNoetherian A K L C) +/-- If `L` is a finite separable extension of `K = Frac(A)`, where `A` is +integrally closed and Noetherian, the integral closure `C` of `A` in `L` is +finite over `A`. -/ +theorem IsIntegralClosure.finite [IsIntegrallyClosed A] [IsNoetherianRing A] : + Module.Finite A C := by + haveI := IsIntegralClosure.isNoetherian A K L C + exact Module.IsNoetherian.finite A C + /-- If `L` is a finite separable extension of `K = Frac(A)`, where `A` is a principal ring and `L` has no zero smul divisors by `A`, the integral closure `C` of `A` in `L` is a free `A`-module. -/ From 76b6113527563ce69dceb7d0d8f56bc2257843cf Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 23 Nov 2024 13:41:10 +0000 Subject: [PATCH 238/829] feat(Ergodic): add `zero_measure` constructors (#19045) --- Mathlib/Dynamics/Ergodic/Ergodic.lean | 13 +++++++++++++ 1 file changed, 13 insertions(+) diff --git a/Mathlib/Dynamics/Ergodic/Ergodic.lean b/Mathlib/Dynamics/Ergodic/Ergodic.lean index b7bd2016eeee4..cc9abb51e7198 100644 --- a/Mathlib/Dynamics/Ergodic/Ergodic.lean +++ b/Mathlib/Dynamics/Ergodic/Ergodic.lean @@ -80,6 +80,9 @@ theorem smul_measure {R : Type*} [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ (hf : PreErgodic f μ) (c : R) : PreErgodic f (c • μ) where aeconst_set _s hs hfs := (hf.aeconst_set hs hfs).anti <| ae_smul_measure_le _ +theorem zero_measure (f : α → α) : @PreErgodic α m f 0 where + aeconst_set _ _ _ := by simp + end PreErgodic namespace MeasureTheory.MeasurePreserving @@ -137,6 +140,11 @@ theorem smul_measure {R : Type*} [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ (hf : QuasiErgodic f μ) (c : R) : QuasiErgodic f (c • μ) := ⟨hf.1.smul_measure _, hf.2.smul_measure _⟩ +theorem zero_measure {f : α → α} (hf : Measurable f) : @QuasiErgodic α m f 0 where + measurable := hf + absolutelyContinuous := by simp + toPreErgodic := .zero_measure f + end QuasiErgodic namespace Ergodic @@ -171,6 +179,11 @@ theorem smul_measure {R : Type*} [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ (hf : Ergodic f μ) (c : R) : Ergodic f (c • μ) := ⟨hf.1.smul_measure _, hf.2.smul_measure _⟩ +theorem zero_measure {f : α → α} (hf : Measurable f) : @Ergodic α m f 0 where + measurable := hf + map_eq := by simp + toPreErgodic := .zero_measure f + section IsFiniteMeasure variable [IsFiniteMeasure μ] From b738c674ff23ff46cb9549010eb25b3a4c1fc79c Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Sat, 23 Nov 2024 13:41:11 +0000 Subject: [PATCH 239/829] =?UTF-8?q?feat(Data/Nat/Nth):=20`nth=20p=20n=20?= =?UTF-8?q?=E2=89=A0=200`=20lemmas=20(#19142)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add a few small lemmas that are convenient when working with `Nat.nth`, when you have a hypothesis of the form `nth p n ≠ 0`. ```lean lemma lt_toFinset_card_of_nth_ne_zero {n : ℕ} (h : nth p n ≠ 0) (hf : (setOf p).Finite) : n < hf.toFinset.card := by ``` ```lean lemma nth_mem_of_ne_zero {n : ℕ} (h : nth p n ≠ 0) : p (Nat.nth p n) := ``` ```lean lemma nth_ne_zero_anti (h₀ : ¬p 0) {a b : ℕ} (hab : a ≤ b) (hb : nth p b ≠ 0) : nth p a ≠ 0 := ``` --- Mathlib/Data/Nat/Nth.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/Mathlib/Data/Nat/Nth.lean b/Mathlib/Data/Nat/Nth.lean index 4cde210602f99..243d52d17d6f6 100644 --- a/Mathlib/Data/Nat/Nth.lean +++ b/Mathlib/Data/Nat/Nth.lean @@ -237,10 +237,21 @@ theorem nth_eq_zero {n} : · rintro (⟨h₀, rfl⟩ | ⟨hf, hle⟩) exacts [nth_zero_of_zero h₀, nth_of_card_le hf hle] +lemma lt_card_toFinset_of_nth_ne_zero {n : ℕ} (h : nth p n ≠ 0) (hf : (setOf p).Finite) : + n < #hf.toFinset := by + simp only [ne_eq, nth_eq_zero, not_or, not_exists, not_le] at h + exact h.2 hf + +lemma nth_mem_of_ne_zero {n : ℕ} (h : nth p n ≠ 0) : p (Nat.nth p n) := + nth_mem n (lt_card_toFinset_of_nth_ne_zero h) + theorem nth_eq_zero_mono (h₀ : ¬p 0) {a b : ℕ} (hab : a ≤ b) (ha : nth p a = 0) : nth p b = 0 := by simp only [nth_eq_zero, h₀, false_and, false_or] at ha ⊢ exact ha.imp fun hf hle => hle.trans hab +lemma nth_ne_zero_anti (h₀ : ¬p 0) {a b : ℕ} (hab : a ≤ b) (hb : nth p b ≠ 0) : nth p a ≠ 0 := + mt (nth_eq_zero_mono h₀ hab) hb + theorem le_nth_of_lt_nth_succ {k a : ℕ} (h : a < nth p (k + 1)) (ha : p a) : a ≤ nth p k := by cases' (setOf p).finite_or_infinite with hf hf · rcases exists_lt_card_finite_nth_eq hf ha with ⟨n, hn, rfl⟩ From a4de385a76ac2e8eeb7aa73f47775fdadf38b287 Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Sat, 23 Nov 2024 13:41:12 +0000 Subject: [PATCH 240/829] feat(CategoryTheory): add some API for countable products and countable filtered colimits (#19192) --- .../Limits/Shapes/Countable.lean | 114 +++++++++++++++++- .../Category/LightProfinite/Basic.lean | 9 +- 2 files changed, 118 insertions(+), 5 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Countable.lean b/Mathlib/CategoryTheory/Limits/Shapes/Countable.lean index a50342b46d933..6f0db0a360ddb 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Countable.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Countable.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Dagur Asgeirsson -/ import Mathlib.CategoryTheory.Limits.Final -import Mathlib.CategoryTheory.Limits.Shapes.FiniteLimits +import Mathlib.CategoryTheory.Limits.Shapes.FiniteProducts import Mathlib.CategoryTheory.Countable import Mathlib.Data.Countable.Defs /-! @@ -52,6 +52,29 @@ instance [Category.{v} J] [CountableCategory J] [HasCountableLimits C] : HasLimi have : HasLimitsOfShape (HomAsType J) C := HasCountableLimits.out (HomAsType J) hasLimitsOfShape_of_equivalence (homAsTypeEquiv J) +/-- A category has countable products if it has all products indexed by countable types. -/ +class HasCountableProducts where + out (J : Type) [Countable J] : HasProductsOfShape J C + +instance [HasCountableProducts C] : HasProductsOfShape J C := + have : Countable (Shrink.{0} J) := Countable.of_equiv _ (equivShrink.{0} J) + have : HasLimitsOfShape (Discrete (Shrink.{0} J)) C := HasCountableProducts.out _ + hasLimitsOfShape_of_equivalence (Discrete.equivalence (equivShrink.{0} J)).symm + +instance (priority := 100) hasCountableProducts_of_hasProducts [HasProducts C] : + HasCountableProducts C where + out _ := + have : HasProducts.{0} C := has_smallest_products_of_hasProducts + inferInstance + +instance (priority := 100) hasCountableProducts_of_hasCountableLimits [HasCountableLimits C] : + HasCountableProducts C where + out _ := inferInstance + +instance (priority := 100) hasFiniteProducts_of_hasCountableProducts [HasCountableProducts C] : + HasFiniteProducts C where + out _ := inferInstance + /-- A category has all countable colimits if every functor `J ⥤ C` with a `CountableCategory J` instance and `J : Type` has a colimit. @@ -74,8 +97,81 @@ instance [Category.{v} J] [CountableCategory J] [HasCountableColimits C] : HasCo have : HasColimitsOfShape (HomAsType J) C := HasCountableColimits.out (HomAsType J) hasColimitsOfShape_of_equivalence (homAsTypeEquiv J) +/-- A category has countable coproducts if it has all coproducts indexed by countable types. -/ +class HasCountableCoproducts where + out (J : Type) [Countable J] : HasCoproductsOfShape J C + +instance (priority := 100) hasCountableCoproducts_of_hasCoproducts [HasCoproducts C] : + HasCountableCoproducts C where + out _ := + have : HasCoproducts.{0} C := has_smallest_coproducts_of_hasCoproducts + inferInstance + +instance [HasCountableCoproducts C] : HasCoproductsOfShape J C := + have : Countable (Shrink.{0} J) := Countable.of_equiv _ (equivShrink.{0} J) + have : HasColimitsOfShape (Discrete (Shrink.{0} J)) C := HasCountableCoproducts.out _ + hasColimitsOfShape_of_equivalence (Discrete.equivalence (equivShrink.{0} J)).symm + +instance (priority := 100) hasCountableCoproducts_of_hasCountableColimits [HasCountableColimits C] : + HasCountableCoproducts C where + out _ := inferInstance + +instance (priority := 100) hasFiniteCoproducts_of_hasCountableCoproducts + [HasCountableCoproducts C] : HasFiniteCoproducts C where + out _ := inferInstance + section Preorder +namespace IsFiltered + +attribute [local instance] IsFiltered.nonempty + +variable {C} [Preorder J] [IsFiltered J] + +/-- The object part of the initial functor `ℕᵒᵖ ⥤ J` -/ +noncomputable def sequentialFunctor_obj : ℕ → J := fun + | .zero => (exists_surjective_nat _).choose 0 + | .succ n => (IsFilteredOrEmpty.cocone_objs ((exists_surjective_nat _).choose n) + (sequentialFunctor_obj n)).choose + +theorem sequentialFunctor_map : Monotone (sequentialFunctor_obj J) := + monotone_nat_of_le_succ fun n ↦ + leOfHom (IsFilteredOrEmpty.cocone_objs ((exists_surjective_nat _).choose n) + (sequentialFunctor_obj J n)).choose_spec.choose_spec.choose + +/-- +The initial functor `ℕᵒᵖ ⥤ J`, which allows us to turn cofiltered limits over countable preorders +into sequential limits. +-/ +noncomputable def sequentialFunctor : ℕ ⥤ J where + obj n := sequentialFunctor_obj J n + map h := homOfLE (sequentialFunctor_map J (leOfHom h)) + +theorem sequentialFunctor_final_aux (j : J) : ∃ (n : ℕ), j ≤ sequentialFunctor_obj J n := by + obtain ⟨m, h⟩ := (exists_surjective_nat _).choose_spec j + refine ⟨m + 1, ?_⟩ + simpa only [h] using leOfHom (IsFilteredOrEmpty.cocone_objs ((exists_surjective_nat _).choose m) + (sequentialFunctor_obj J m)).choose_spec.choose + +instance sequentialFunctor_final : (sequentialFunctor J).Final where + out d := by + obtain ⟨n, (g : d ≤ (sequentialFunctor J).obj n)⟩ := sequentialFunctor_final_aux J d + have : Nonempty (StructuredArrow d (sequentialFunctor J)) := + ⟨StructuredArrow.mk (homOfLE g)⟩ + apply isConnected_of_zigzag + refine fun i j ↦ ⟨[j], ?_⟩ + simp only [List.chain_cons, Zag, List.Chain.nil, and_true, ne_eq, not_false_eq_true, + List.getLast_cons, not_true_eq_false, List.getLast_singleton', reduceCtorEq] + clear! C + wlog h : j.right ≤ i.right + · exact or_comm.1 (this J d n g inferInstance j i (le_of_lt (not_le.mp h))) + · right + exact ⟨StructuredArrow.homMk (homOfLE h) rfl⟩ + +end IsFiltered + +namespace IsCofiltered + attribute [local instance] IsCofiltered.nonempty variable {C} [Preorder J] [IsCofiltered J] @@ -94,6 +190,9 @@ theorem sequentialFunctor_map : Antitone (sequentialFunctor_obj J) := /-- The initial functor `ℕᵒᵖ ⥤ J`, which allows us to turn cofiltered limits over countable preorders into sequential limits. + +TODO: redefine this as `(IsFiltered.sequentialFunctor Jᵒᵖ).leftOp`. This would need API for initial/ +final functors of the form `leftOp`/`rightOp`. -/ noncomputable def sequentialFunctor : ℕᵒᵖ ⥤ J where obj n := sequentialFunctor_obj J (unop n) @@ -153,6 +252,19 @@ For this we need to dualize this whole section. proof_wanted hasCountableColimits_of_hasFiniteColimits_and_hasSequentialColimits [HasFiniteColimits C] [HasLimitsOfShape ℕ C] : HasCountableColimits C +end IsCofiltered + end Preorder +@[deprecated (since := "2024-11-01")] alias sequentialFunctor := IsCofiltered.sequentialFunctor +@[deprecated (since := "2024-11-01")] alias sequentialFunctor_obj := + IsCofiltered.sequentialFunctor_obj +@[deprecated (since := "2024-11-01")] alias sequentialFunctor_map := + IsCofiltered.sequentialFunctor_map +@[deprecated (since := "2024-11-01")] alias sequentialFunctor_initial_aux := + IsCofiltered.sequentialFunctor_initial_aux +@[deprecated (since := "2024-11-01")] alias sequentialFunctor_initial := + IsCofiltered.sequentialFunctor_initial +attribute [nolint defLemma] sequentialFunctor_initial + end CategoryTheory.Limits diff --git a/Mathlib/Topology/Category/LightProfinite/Basic.lean b/Mathlib/Topology/Category/LightProfinite/Basic.lean index 362a23b4df054..bd67bbffa19ea 100644 --- a/Mathlib/Topology/Category/LightProfinite/Basic.lean +++ b/Mathlib/Topology/Category/LightProfinite/Basic.lean @@ -297,10 +297,11 @@ instance instCountableDiscreteQuotient (S : LightProfinite) : /-- A profinite space which is light gives rise to a light profinite space. -/ noncomputable def toLightDiagram (S : LightProfinite.{u}) : LightDiagram.{u} where - diagram := sequentialFunctor _ ⋙ (lightToProfinite.obj S).fintypeDiagram - cone := (Functor.Initial.limitConeComp (sequentialFunctor _) (lightToProfinite.obj S).lim).cone - isLimit := - (Functor.Initial.limitConeComp (sequentialFunctor _) (lightToProfinite.obj S).lim).isLimit + diagram := IsCofiltered.sequentialFunctor _ ⋙ (lightToProfinite.obj S).fintypeDiagram + cone := (Functor.Initial.limitConeComp (IsCofiltered.sequentialFunctor _) + (lightToProfinite.obj S).lim).cone + isLimit := (Functor.Initial.limitConeComp (IsCofiltered.sequentialFunctor _) + (lightToProfinite.obj S).lim).isLimit end LightProfinite From 1b79c7d8d77de09fa646c73ea18316f7d4fe6575 Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Sat, 23 Nov 2024 13:41:13 +0000 Subject: [PATCH 241/829] feat(CategoryTheory): add `Pi.map_isIso` and `hasProductsOfShape_of_hasProducts` + duals (#19195) --- Mathlib/CategoryTheory/Limits/Shapes/Products.lean | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Products.lean b/Mathlib/CategoryTheory/Limits/Shapes/Products.lean index 0bf114008f136..ed9bd8a77cbda 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Products.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Products.lean @@ -373,6 +373,10 @@ from a family of isomorphisms between the factors. abbrev Pi.mapIso {f g : β → C} [HasProductsOfShape β C] (p : ∀ b, f b ≅ g b) : ∏ᶜ f ≅ ∏ᶜ g := lim.mapIso (Discrete.natIso fun X => p X.as) +instance Pi.map_isIso {f g : β → C} [HasProductsOfShape β C] (p : ∀ b, f b ⟶ g b) + [∀ b, IsIso <| p b] : IsIso <| Pi.map p := + inferInstanceAs (IsIso (Pi.mapIso (fun b ↦ asIso (p b))).hom) + section /- In this section, we provide some API for products when we are given a functor @@ -487,6 +491,10 @@ from a family of isomorphisms between the factors. abbrev Sigma.mapIso {f g : β → C} [HasCoproductsOfShape β C] (p : ∀ b, f b ≅ g b) : ∐ f ≅ ∐ g := colim.mapIso (Discrete.natIso fun X => p X.as) +instance Sigma.map_isIso {f g : β → C} [HasCoproductsOfShape β C] (p : ∀ b, f b ⟶ g b) + [∀ b, IsIso <| p b] : IsIso (Sigma.map p) := + inferInstanceAs (IsIso (Sigma.mapIso (fun b ↦ asIso (p b))).hom) + section /- In this section, we provide some API for coproducts when we are given a functor @@ -660,6 +668,12 @@ theorem hasProducts_of_limit_fans (lf : ∀ {J : Type w} (f : J → C), Fan f) ⟨(Cones.postcompose Discrete.natIsoFunctor.inv).obj (lf fun j => F.obj ⟨j⟩), (IsLimit.postcomposeInvEquiv _ _).symm (lf_isLimit _)⟩ } +instance (priority := 100) hasProductsOfShape_of_hasProducts [HasProducts.{w} C] (J : Type w) : + HasProductsOfShape J C := inferInstance + +instance (priority := 100) hasCoproductsOfShape_of_hasCoproducts [HasCoproducts.{w} C] + (J : Type w) : HasCoproductsOfShape J C := inferInstance + /-! (Co)products over a type with a unique term. -/ From 6f5a5765040cc6b1b7c36c3f80821448977a8899 Mon Sep 17 00:00:00 2001 From: damiano Date: Sat, 23 Nov 2024 13:41:14 +0000 Subject: [PATCH 242/829] feat: add "new file" markers to import reports (#19274) Flagging new files should help in identifying why some files appear to have increased the number of imports, when creating new files. This was suggested in #19194. --- scripts/import_trans_difference.sh | 14 ++++++++++++-- 1 file changed, 12 insertions(+), 2 deletions(-) diff --git a/scripts/import_trans_difference.sh b/scripts/import_trans_difference.sh index 9c103e3eabdae..a13bebc4957cf 100755 --- a/scripts/import_trans_difference.sh +++ b/scripts/import_trans_difference.sh @@ -67,7 +67,16 @@ git checkout "${currCommit}" printf '\n\n
Import changes for all files\n\n%s\n\n
\n' "$( printf "|Files|Import difference|\n|-|-|\n" - (awk -F, -v all="${all}" -v ghLimit='261752' '{ diff[$1]+=$2 } END { + (awk -F, -v all="${all}" -v ghLimit='261752' -v newFiles="$( + # we pass the "A"dded files with respect to master, converting them to module names + git diff --name-only --diff-filter=A master | tr '\n' , | sed 's=\.lean,=,=g; s=/=.=g' + )" ' + BEGIN{ + # `arrayNewModules` maps integers to module names + split(newFiles, arrayNewModules, ",") + # `newModules` "just" stores the module names + for(v in arrayNewModules) { newModules[arrayNewModules[v]]=0 } + } { diff[$1]+=$2 } END { fileCount=0 outputLength=0 for(fil in diff) { @@ -75,7 +84,8 @@ printf '\n\n
Import changes for all files\n\n%s\n\n Date: Sat, 23 Nov 2024 13:41:16 +0000 Subject: [PATCH 243/829] chore: rename *of_smul_regular lemmas (#19326) --- Mathlib/LinearAlgebra/Dimension/Basic.lean | 5 ++++- Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean | 9 ++++++--- .../LinearAlgebra/RootSystem/Finite/Nondegenerate.lean | 2 +- 3 files changed, 11 insertions(+), 5 deletions(-) diff --git a/Mathlib/LinearAlgebra/Dimension/Basic.lean b/Mathlib/LinearAlgebra/Dimension/Basic.lean index 254ddee6fae30..16bbd326c3112 100644 --- a/Mathlib/LinearAlgebra/Dimension/Basic.lean +++ b/Mathlib/LinearAlgebra/Dimension/Basic.lean @@ -353,13 +353,16 @@ theorem rank_subsingleton [Subsingleton R] : Module.rank R M = 1 := by subsingleton · exact hw.trans_eq (Cardinal.mk_singleton _).symm -lemma rank_le_of_smul_regular {S : Type*} [CommSemiring S] [Algebra S R] [Module S M] +lemma rank_le_of_isSMulRegular {S : Type*} [CommSemiring S] [Algebra S R] [Module S M] [IsScalarTower S R M] (L L' : Submodule R M) {s : S} (hr : IsSMulRegular M s) (h : ∀ x ∈ L, s • x ∈ L') : Module.rank R L ≤ Module.rank R L' := ((Algebra.lsmul S R M s).restrict h).rank_le_of_injective <| fun _ _ h ↦ by simpa using hr (Subtype.ext_iff.mp h) +@[deprecated (since := "2024-11-21")] +alias rank_le_of_smul_regular := rank_le_of_isSMulRegular + end end Module diff --git a/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean b/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean index 0e7836ca78bb4..71be346b20eaa 100644 --- a/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean +++ b/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean @@ -465,12 +465,15 @@ theorem LinearMap.finrank_range_le [Module.Finite R M] (f : M →ₗ[R] M') : finrank R (LinearMap.range f) ≤ finrank R M := finrank_le_finrank_of_rank_le_rank (lift_rank_range_le f) (rank_lt_aleph0 _ _) -theorem LinearMap.finrank_le_of_smul_regular {S : Type*} [CommSemiring S] [Algebra S R] [Module S M] - [IsScalarTower S R M] (L L' : Submodule R M) [Module.Finite R L'] {s : S} +theorem LinearMap.finrank_le_of_isSMulRegular {S : Type*} [CommSemiring S] [Algebra S R] + [Module S M] [IsScalarTower S R M] (L L' : Submodule R M) [Module.Finite R L'] {s : S} (hr : IsSMulRegular M s) (h : ∀ x ∈ L, s • x ∈ L') : Module.finrank R L ≤ Module.finrank R L' := by - refine finrank_le_finrank_of_rank_le_rank (lift_le.mpr <| rank_le_of_smul_regular L L' hr h) ?_ + refine finrank_le_finrank_of_rank_le_rank (lift_le.mpr <| rank_le_of_isSMulRegular L L' hr h) ?_ rw [← Module.finrank_eq_rank R L'] exact nat_lt_aleph0 (finrank R ↥L') +@[deprecated (since := "2024-11-21")] +alias LinearMap.finrank_le_of_smul_regular := LinearMap.finrank_le_of_isSMulRegular + end StrongRankCondition diff --git a/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean b/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean index 734860e88627e..1efe86c29f24f 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean @@ -70,7 +70,7 @@ lemma finrank_rootSpan_map_polarization_eq_finrank_corootSpan : rw [← LinearMap.range_domRestrict] apply (Submodule.finrank_mono P.range_polarization_domRestrict_le_span_coroot).antisymm have : IsReflexive R N := PerfectPairing.reflexive_right P.toPerfectPairing - refine LinearMap.finrank_le_of_smul_regular P.corootSpan + refine LinearMap.finrank_le_of_isSMulRegular P.corootSpan (LinearMap.range (P.Polarization.domRestrict P.rootSpan)) (smul_right_injective N (Ne.symm (ne_of_lt P.prod_rootForm_root_self_pos))) fun _ hx => ?_ From d6a85bb4f0e582e6f48ee7c49c1c6ce8601797f7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 23 Nov 2024 14:14:52 +0000 Subject: [PATCH 244/829] feat(Polynomial): `p.natDegree = q.natDegree` if `p.degree = q.degree` (#19402) Also generalise `natDegree_lt_natDegree` to two rings From GrowthInGroups (LeanCamCombi) Co-authored-by: Andrew Yang --- Mathlib/Algebra/Polynomial/Degree/Operations.lean | 7 +++++-- 1 file changed, 5 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/Degree/Operations.lean b/Mathlib/Algebra/Polynomial/Degree/Operations.lean index b64a85909e805..21d80bcb9d741 100644 --- a/Mathlib/Algebra/Polynomial/Degree/Operations.lean +++ b/Mathlib/Algebra/Polynomial/Degree/Operations.lean @@ -29,7 +29,7 @@ variable {R : Type u} {S : Type v} {a b c d : R} {n m : ℕ} section Semiring -variable [Semiring R] {p q r : R[X]} +variable [Semiring R] [Semiring S] {p q r : R[X]} theorem supDegree_eq_degree (p : R[X]) : p.toFinsupp.supDegree WithBot.some = p.degree := max_eq_sup_coe @@ -65,12 +65,15 @@ theorem natDegree_eq_of_le_of_coeff_ne_zero (pn : p.natDegree ≤ n) (p1 : p.coe p.natDegree = n := pn.antisymm (le_natDegree_of_ne_zero p1) -theorem natDegree_lt_natDegree {p q : R[X]} (hp : p ≠ 0) (hpq : p.degree < q.degree) : +theorem natDegree_lt_natDegree {q : S[X]} (hp : p ≠ 0) (hpq : p.degree < q.degree) : p.natDegree < q.natDegree := by by_cases hq : q = 0 · exact (not_lt_bot <| hq ▸ hpq).elim rwa [degree_eq_natDegree hp, degree_eq_natDegree hq, Nat.cast_lt] at hpq +lemma natDegree_eq_natDegree {q : S[X]} (hpq : p.degree = q.degree) : + p.natDegree = q.natDegree := by simp [natDegree, hpq] + theorem coeff_eq_zero_of_degree_lt (h : degree p < n) : coeff p n = 0 := Classical.not_not.1 (mt le_degree_of_ne_zero (not_le_of_gt h)) From 2abe270b9a11bad57afac01e29a40b4e92e1fcc5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 23 Nov 2024 14:14:53 +0000 Subject: [PATCH 245/829] chore: rename `Set.nonempty_of_nonempty_subtype` to `Set.Nonempty.of_subtype` (#19408) This is shorter and allows anonymous dot notation. From GrowthInGroups (LeanCamCombi) --- Mathlib/Data/ENat/Lattice.lean | 2 +- Mathlib/Data/Set/Basic.lean | 7 ++++--- Mathlib/Data/Set/Image.lean | 2 +- Mathlib/GroupTheory/PGroup.lean | 2 +- 4 files changed, 7 insertions(+), 6 deletions(-) diff --git a/Mathlib/Data/ENat/Lattice.lean b/Mathlib/Data/ENat/Lattice.lean index d56040d7d9e21..fd0d60e48e317 100644 --- a/Mathlib/Data/ENat/Lattice.lean +++ b/Mathlib/Data/ENat/Lattice.lean @@ -106,7 +106,7 @@ lemma finite_of_sSup_lt_top (h : sSup s < ⊤) : s.Finite := by exact sSup_eq_top_of_infinite h lemma sSup_mem_of_nonempty_of_lt_top [Nonempty s] (hs' : sSup s < ⊤) : sSup s ∈ s := - Nonempty.csSup_mem nonempty_of_nonempty_subtype (finite_of_sSup_lt_top hs') + Nonempty.csSup_mem .of_subtype (finite_of_sSup_lt_top hs') lemma exists_eq_iSup_of_lt_top [Nonempty ι] (h : ⨆ i, f i < ⊤) : ∃ i, f i = ⨆ i, f i := diff --git a/Mathlib/Data/Set/Basic.lean b/Mathlib/Data/Set/Basic.lean index d5e0c9c761f33..8d28bac27f0b0 100644 --- a/Mathlib/Data/Set/Basic.lean +++ b/Mathlib/Data/Set/Basic.lean @@ -436,8 +436,9 @@ instance univ.nonempty [Nonempty α] : Nonempty (↥(Set.univ : Set α)) := instance instNonemptyTop [Nonempty α] : Nonempty (⊤ : Set α) := inferInstanceAs (Nonempty (univ : Set α)) -theorem nonempty_of_nonempty_subtype [Nonempty (↥s)] : s.Nonempty := - nonempty_subtype.mp ‹_› +theorem Nonempty.of_subtype [Nonempty (↥s)] : s.Nonempty := nonempty_subtype.mp ‹_› + +@[deprecated (since := "2024-11-23")] alias nonempty_of_nonempty_subtype := Nonempty.of_subtype /-! ### Lemmas about the empty set -/ @@ -1204,7 +1205,7 @@ theorem eq_empty_of_ssubset_singleton {s : Set α} {x : α} (hs : s ⊂ {x}) : s theorem eq_of_nonempty_of_subsingleton {α} [Subsingleton α] (s t : Set α) [Nonempty s] [Nonempty t] : s = t := - nonempty_of_nonempty_subtype.eq_univ.trans nonempty_of_nonempty_subtype.eq_univ.symm + Nonempty.of_subtype.eq_univ.trans Nonempty.of_subtype.eq_univ.symm theorem eq_of_nonempty_of_subsingleton' {α} [Subsingleton α] {s : Set α} (t : Set α) (hs : s.Nonempty) [Nonempty t] : s = t := diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index 11aa60b7641a2..16c6dcc43e64f 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -408,7 +408,7 @@ theorem Nonempty.preimage {s : Set β} (hs : s.Nonempty) {f : α → β} (hf : S ⟨x, mem_preimage.2 <| hx.symm ▸ hy⟩ instance (f : α → β) (s : Set α) [Nonempty s] : Nonempty (f '' s) := - (Set.Nonempty.image f nonempty_of_nonempty_subtype).to_subtype + (Set.Nonempty.image f .of_subtype).to_subtype /-- image and preimage are a Galois connection -/ @[simp] diff --git a/Mathlib/GroupTheory/PGroup.lean b/Mathlib/GroupTheory/PGroup.lean index 9b119d28f05df..071291b214ad5 100644 --- a/Mathlib/GroupTheory/PGroup.lean +++ b/Mathlib/GroupTheory/PGroup.lean @@ -189,7 +189,7 @@ theorem card_modEq_card_fixedPoints : Nat.card α ≡ Nat.card (fixedPoints G α theorem nonempty_fixed_point_of_prime_not_dvd_card (α) [MulAction G α] (hpα : ¬p ∣ Nat.card α) : (fixedPoints G α).Nonempty := have : Finite α := Nat.finite_of_card_ne_zero (fun h ↦ (h ▸ hpα) (dvd_zero p)) - @Set.nonempty_of_nonempty_subtype _ _ + @Set.Nonempty.of_subtype _ _ (by rw [← Finite.card_pos_iff, pos_iff_ne_zero] contrapose! hpα From 1620b76241d3341c82b959798ef4b6e44e1dc626 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Sat, 23 Nov 2024 14:58:15 +0000 Subject: [PATCH 246/829] feat(AlgebraicGeometry): preimmersions are stable under base change (#18915) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib.lean | 1 + .../Cover/MorphismProperty.lean | 3 + .../Morphisms/ClosedImmersion.lean | 4 +- .../Morphisms/Preimmersion.lean | 40 ++-- .../Morphisms/SurjectiveOnStalks.lean | 186 ++++++++++++++++++ Mathlib/AlgebraicGeometry/Scheme.lean | 12 ++ Mathlib/Topology/LocalAtTarget.lean | 41 ++++ 7 files changed, 264 insertions(+), 23 deletions(-) create mode 100644 Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean diff --git a/Mathlib.lean b/Mathlib.lean index 48be73827fdd6..58d4cfdeb00c9 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -962,6 +962,7 @@ import Mathlib.AlgebraicGeometry.Morphisms.QuasiSeparated import Mathlib.AlgebraicGeometry.Morphisms.RingHomProperties import Mathlib.AlgebraicGeometry.Morphisms.Separated import Mathlib.AlgebraicGeometry.Morphisms.Smooth +import Mathlib.AlgebraicGeometry.Morphisms.SurjectiveOnStalks import Mathlib.AlgebraicGeometry.Morphisms.UnderlyingMap import Mathlib.AlgebraicGeometry.Morphisms.UniversallyClosed import Mathlib.AlgebraicGeometry.Morphisms.UniversallyInjective diff --git a/Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean b/Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean index e2226dd86717f..53a9c4dfd2092 100644 --- a/Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean +++ b/Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean @@ -73,6 +73,9 @@ theorem Cover.iUnion_range {X : Scheme.{u}} (𝒰 : X.Cover P) : rw [Set.mem_iUnion] exact ⟨𝒰.f x, 𝒰.covers x⟩ +lemma Cover.exists_eq (𝒰 : X.Cover P) (x : X) : ∃ i y, (𝒰.map i).base y = x := + ⟨_, 𝒰.covers x⟩ + /-- Given a family of schemes with morphisms to `X` satisfying `P` that jointly cover `X`, this an associated `P`-cover of `X`. -/ @[simps] diff --git a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean index 1cb1635d9fcb0..d72195178231c 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean @@ -62,8 +62,8 @@ lemma eq_inf : @IsClosedImmersion = (topologically IsClosedEmbedding) ⊓ lemma iff_isPreimmersion {X Y : Scheme} {f : X ⟶ Y} : IsClosedImmersion f ↔ IsPreimmersion f ∧ IsClosed (Set.range f.base) := by - rw [and_comm, isClosedImmersion_iff, isPreimmersion_iff, ← and_assoc, isClosedEmbedding_iff, - @and_comm (IsEmbedding _)] + rw [isClosedImmersion_iff, isPreimmersion_iff, ← surjectiveOnStalks_iff, and_comm, and_assoc, + isClosedEmbedding_iff] lemma of_isPreimmersion {X Y : Scheme} (f : X ⟶ Y) [IsPreimmersion f] (hf : IsClosed (Set.range f.base)) : IsClosedImmersion f := diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean index 4d5503fe1dc7f..a2505954c2c6e 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean @@ -4,8 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ import Mathlib.AlgebraicGeometry.Morphisms.UnderlyingMap -import Mathlib.RingTheory.RingHom.Surjective -import Mathlib.RingTheory.SurjectiveOnStalks +import Mathlib.AlgebraicGeometry.Morphisms.SurjectiveOnStalks /-! @@ -33,9 +32,8 @@ namespace AlgebraicGeometry /-- A morphism of schemes `f : X ⟶ Y` is a preimmersion if the underlying map of topological spaces is an embedding and the induced morphisms of stalks are all surjective. -/ @[mk_iff] -class IsPreimmersion {X Y : Scheme} (f : X ⟶ Y) : Prop where +class IsPreimmersion {X Y : Scheme} (f : X ⟶ Y) extends SurjectiveOnStalks f : Prop where base_embedding : IsEmbedding f.base - surj_on_stalks : ∀ x, Function.Surjective (f.stalkMap x) lemma Scheme.Hom.isEmbedding {X Y : Scheme} (f : Hom X Y) [IsPreimmersion f] : IsEmbedding f.base := IsPreimmersion.base_embedding @@ -43,12 +41,8 @@ lemma Scheme.Hom.isEmbedding {X Y : Scheme} (f : Hom X Y) [IsPreimmersion f] : I @[deprecated (since := "2024-10-26")] alias Scheme.Hom.embedding := Scheme.Hom.isEmbedding -lemma Scheme.Hom.stalkMap_surjective {X Y : Scheme} (f : Hom X Y) [IsPreimmersion f] (x) : - Function.Surjective (f.stalkMap x) := - IsPreimmersion.surj_on_stalks x - lemma isPreimmersion_eq_inf : - @IsPreimmersion = topologically IsEmbedding ⊓ stalkwise (Function.Surjective ·) := by + @IsPreimmersion = (@SurjectiveOnStalks ⊓ topologically IsEmbedding : MorphismProperty _) := by ext rw [isPreimmersion_iff] rfl @@ -69,10 +63,7 @@ instance (priority := 900) {X Y : Scheme} (f : X ⟶ Y) [IsOpenImmersion f] : Is instance : MorphismProperty.IsMultiplicative @IsPreimmersion where id_mem _ := inferInstance - comp_mem {X Y Z} f g hf hg := by - refine ⟨hg.base_embedding.comp hf.base_embedding, fun x ↦ ?_⟩ - rw [Scheme.stalkMap_comp] - exact (hf.surj_on_stalks x).comp (hg.surj_on_stalks (f.base x)) + comp_mem f g _ _ := ⟨g.isEmbedding.comp f.isEmbedding⟩ instance comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [IsPreimmersion f] [IsPreimmersion g] : IsPreimmersion (f ≫ g) := @@ -104,14 +95,8 @@ lemma Spec_map_iff {R S : CommRingCat.{u}} (f : R ⟶ S) : haveI : (RingHom.toMorphismProperty <| fun f ↦ Function.Surjective f).RespectsIso := by rw [← RingHom.toMorphismProperty_respectsIso_iff] exact RingHom.surjective_respectsIso - refine ⟨fun ⟨h₁, h₂⟩ ↦ ⟨h₁, ?_⟩, fun ⟨h₁, h₂⟩ ↦ ⟨h₁, fun (x : PrimeSpectrum S) ↦ ?_⟩⟩ - · intro p hp - let e := Scheme.arrowStalkMapSpecIso f ⟨p, hp⟩ - apply ((RingHom.toMorphismProperty <| fun f ↦ Function.Surjective f).arrow_mk_iso_iff e).mp - exact h₂ ⟨p, hp⟩ - · let e := Scheme.arrowStalkMapSpecIso f x - apply ((RingHom.toMorphismProperty <| fun f ↦ Function.Surjective f).arrow_mk_iso_iff e).mpr - exact h₂ x.asIdeal x.isPrime + rw [← HasRingHomProperty.Spec_iff (P := @SurjectiveOnStalks), isPreimmersion_iff, and_comm] + rfl lemma mk_Spec_map {R S : CommRingCat.{u}} {f : R ⟶ S} (h₁ : IsEmbedding (PrimeSpectrum.comap f)) (h₂ : f.SurjectiveOnStalks) : @@ -125,6 +110,19 @@ lemma of_isLocalization {R S : Type u} [CommRing R] (M : Submonoid R) [CommRing (PrimeSpectrum.localization_comap_isEmbedding (R := R) S M) (RingHom.surjectiveOnStalks_of_isLocalization (M := M) S) +open Limits MorphismProperty in +instance : IsStableUnderBaseChange @IsPreimmersion := by + refine .mk' fun X Y Z f g _ _ ↦ ?_ + have := pullback_fst (P := @SurjectiveOnStalks) f g inferInstance + constructor + let L (x : (pullback f g : _)) : { x : X × Y | f.base x.1 = g.base x.2 } := + ⟨⟨(pullback.fst f g).base x, (pullback.snd f g).base x⟩, + by simp only [Set.mem_setOf, ← Scheme.comp_base_apply, pullback.condition]⟩ + have : IsEmbedding L := IsEmbedding.of_comp (by fun_prop) continuous_subtype_val + (SurjectiveOnStalks.isEmbedding_pullback f g) + exact IsEmbedding.subtypeVal.comp ((TopCat.pullbackHomeoPreimage _ f.continuous _ + g.isEmbedding).isEmbedding.comp this) + end IsPreimmersion end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean b/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean new file mode 100644 index 0000000000000..548f463af6395 --- /dev/null +++ b/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean @@ -0,0 +1,186 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.AlgebraicGeometry.Morphisms.RingHomProperties +import Mathlib.AlgebraicGeometry.PrimeSpectrum.TensorProduct +import Mathlib.Topology.LocalAtTarget + +/-! +# Morphisms surjective on stalks + +We define the classe of morphisms between schemes that are surjective on stalks. +We show that this class is stable under composition and base change. + +We also show that (`AlgebraicGeometry.SurjectiveOnStalks.isEmbedding_pullback`) +if `Y ⟶ S` is surjective on stalks, then for every `X ⟶ S`, `X ×ₛ Y` is a subset of +`X × Y` (cartesian product as topological spaces) with the induced topology. +-/ + +open CategoryTheory CategoryTheory.Limits Topology + +namespace AlgebraicGeometry + +universe u + +variable {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) + +/-- The class of morphisms `f : X ⟶ Y` between schemes such that +`𝒪_{Y, f x} ⟶ 𝒪_{X, x}` is surjective for all `x : X`. -/ +@[mk_iff] +class SurjectiveOnStalks : Prop where + surj_on_stalks : ∀ x, Function.Surjective (f.stalkMap x) + +theorem Scheme.Hom.stalkMap_surjective (f : X.Hom Y) [SurjectiveOnStalks f] (x) : + Function.Surjective (f.stalkMap x) := + SurjectiveOnStalks.surj_on_stalks x + +namespace SurjectiveOnStalks + +instance (priority := 900) [IsOpenImmersion f] : SurjectiveOnStalks f := + ⟨fun _ ↦ (ConcreteCategory.bijective_of_isIso _).2⟩ + +instance : MorphismProperty.IsMultiplicative @SurjectiveOnStalks where + id_mem _ := inferInstance + comp_mem {X Y Z} f g hf hg := by + refine ⟨fun x ↦ ?_⟩ + rw [Scheme.stalkMap_comp] + exact (hf.surj_on_stalks x).comp (hg.surj_on_stalks (f.base x)) + +instance comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [SurjectiveOnStalks f] + [SurjectiveOnStalks g] : SurjectiveOnStalks (f ≫ g) := + MorphismProperty.IsStableUnderComposition.comp_mem f g inferInstance inferInstance + +lemma eq_stalkwise : + @SurjectiveOnStalks = stalkwise (Function.Surjective ·) := by + ext; exact surjectiveOnStalks_iff _ + +instance : IsLocalAtTarget @SurjectiveOnStalks := + eq_stalkwise ▸ stalkwiseIsLocalAtTarget_of_respectsIso RingHom.surjective_respectsIso + +instance : IsLocalAtSource @SurjectiveOnStalks := + eq_stalkwise ▸ stalkwise_isLocalAtSource_of_respectsIso RingHom.surjective_respectsIso + +lemma Spec_iff {R S : CommRingCat.{u}} {φ : R ⟶ S} : + SurjectiveOnStalks (Spec.map φ) ↔ RingHom.SurjectiveOnStalks φ := by + rw [eq_stalkwise, stalkwise_Spec_map_iff RingHom.surjective_respectsIso, + RingHom.SurjectiveOnStalks] + +instance : HasRingHomProperty @SurjectiveOnStalks RingHom.SurjectiveOnStalks := + eq_stalkwise ▸ .stalkwise RingHom.surjective_respectsIso + +variable {f} in +lemma iff_of_isAffine [IsAffine X] [IsAffine Y] : + SurjectiveOnStalks f ↔ RingHom.SurjectiveOnStalks (f.app ⊤) := by + rw [← Spec_iff, MorphismProperty.arrow_mk_iso_iff @SurjectiveOnStalks (arrowIsoSpecΓOfIsAffine f)] + rfl + +theorem of_comp [SurjectiveOnStalks (f ≫ g)] : SurjectiveOnStalks f := by + refine ⟨fun x ↦ ?_⟩ + have := (f ≫ g).stalkMap_surjective x + rw [Scheme.stalkMap_comp] at this + exact Function.Surjective.of_comp this + +instance stableUnderBaseChange : + MorphismProperty.IsStableUnderBaseChange @SurjectiveOnStalks := by + apply HasRingHomProperty.isStableUnderBaseChange + apply RingHom.IsStableUnderBaseChange.mk + · exact (HasRingHomProperty.isLocal_ringHomProperty @SurjectiveOnStalks).respectsIso + intros R S T _ _ _ _ _ H + exact H.baseChange + +/-- If `Y ⟶ S` is surjective on stalks, then for every `X ⟶ S`, `X ×ₛ Y` is a subset of +`X × Y` (cartesian product as topological spaces) with the induced topology. -/ +lemma isEmbedding_pullback {X Y S : Scheme.{u}} (f : X ⟶ S) (g : Y ⟶ S) [SurjectiveOnStalks g] : + IsEmbedding (fun x ↦ ((pullback.fst f g).base x, (pullback.snd f g).base x)) := by + let L := (fun x ↦ ((pullback.fst f g).base x, (pullback.snd f g).base x)) + have H : ∀ R A B (f' : Spec A ⟶ Spec R) (g' : Spec B ⟶ Spec R) (iX : Spec A ⟶ X) + (iY : Spec B ⟶ Y) (iS : Spec R ⟶ S) (e₁ e₂), IsOpenImmersion iX → IsOpenImmersion iY → + IsOpenImmersion iS → IsEmbedding (L ∘ (pullback.map f' g' f g iX iY iS e₁ e₂).base) := by + intro R A B f' g' iX iY iS e₁ e₂ _ _ _ + have H : SurjectiveOnStalks g' := + have : SurjectiveOnStalks (g' ≫ iS) := e₂ ▸ inferInstance + .of_comp _ iS + obtain ⟨φ, rfl⟩ : ∃ φ, Spec.map φ = f' := ⟨_, Spec.map_preimage _⟩ + obtain ⟨ψ, rfl⟩ : ∃ ψ, Spec.map ψ = g' := ⟨_, Spec.map_preimage _⟩ + letI := φ.toAlgebra + letI := ψ.toAlgebra + rw [HasRingHomProperty.Spec_iff (P := @SurjectiveOnStalks)] at H + convert ((iX.isOpenEmbedding.prodMap iY.isOpenEmbedding).isEmbedding.comp + (PrimeSpectrum.isEmbedding_tensorProductTo_of_surjectiveOnStalks R A B H)).comp + (Scheme.homeoOfIso (pullbackSpecIso R A B)).isEmbedding + ext1 x + obtain ⟨x, rfl⟩ := (Scheme.homeoOfIso (pullbackSpecIso R A B).symm).surjective x + simp only [Scheme.homeoOfIso_apply, Function.comp_apply] + ext + · simp only [← Scheme.comp_base_apply, pullback.lift_fst, Iso.symm_hom, Iso.inv_hom_id] + erw [← Scheme.comp_base_apply, pullbackSpecIso_inv_fst_assoc] + rfl + · simp only [← Scheme.comp_base_apply, pullback.lift_snd, Iso.symm_hom, Iso.inv_hom_id] + erw [← Scheme.comp_base_apply, pullbackSpecIso_inv_snd_assoc] + rfl + let 𝒰 := S.affineOpenCover.openCover + let 𝒱 (i) := ((𝒰.pullbackCover f).obj i).affineOpenCover.openCover + let 𝒲 (i) := ((𝒰.pullbackCover g).obj i).affineOpenCover.openCover + let U (ijk : Σ i, (𝒱 i).J × (𝒲 i).J) : TopologicalSpace.Opens (X.carrier × Y) := + ⟨{ P | P.1 ∈ ((𝒱 ijk.1).map ijk.2.1 ≫ (𝒰.pullbackCover f).map ijk.1).opensRange ∧ + P.2 ∈ ((𝒲 ijk.1).map ijk.2.2 ≫ (𝒰.pullbackCover g).map ijk.1).opensRange }, + (continuous_fst.1 _ ((𝒱 ijk.1).map ijk.2.1 ≫ + (𝒰.pullbackCover f).map ijk.1).opensRange.2).inter (continuous_snd.1 _ + ((𝒲 ijk.1).map ijk.2.2 ≫ (𝒰.pullbackCover g).map ijk.1).opensRange.2)⟩ + have : Set.range L ⊆ (iSup U : _) := by + simp only [Scheme.Cover.pullbackCover_J, Scheme.Cover.pullbackCover_obj, Set.range_subset_iff] + intro z + simp only [SetLike.mem_coe, TopologicalSpace.Opens.mem_iSup, Sigma.exists, Prod.exists] + obtain ⟨is, s, hsx⟩ := 𝒰.exists_eq (f.base ((pullback.fst f g).base z)) + have hsy : (𝒰.map is).base s = g.base ((pullback.snd f g).base z) := by + rwa [← Scheme.comp_base_apply, ← pullback.condition, Scheme.comp_base_apply] + obtain ⟨x : (𝒰.pullbackCover f).obj is, hx⟩ := + Scheme.IsJointlySurjectivePreserving.exists_preimage_fst_triplet_of_prop + (P := @IsOpenImmersion) inferInstance _ _ hsx.symm + obtain ⟨y : (𝒰.pullbackCover g).obj is, hy⟩ := + Scheme.IsJointlySurjectivePreserving.exists_preimage_fst_triplet_of_prop + (P := @IsOpenImmersion) inferInstance _ _ hsy.symm + obtain ⟨ix, x, rfl⟩ := (𝒱 is).exists_eq x + obtain ⟨iy, y, rfl⟩ := (𝒲 is).exists_eq y + refine ⟨is, ix, iy, ⟨x, hx⟩, ⟨y, hy⟩⟩ + let 𝓤 := (Scheme.Pullback.openCoverOfBase 𝒰 f g).bind + (fun i ↦ Scheme.Pullback.openCoverOfLeftRight (𝒱 i) (𝒲 i) _ _) + refine isEmbedding_of_iSup_eq_top_of_preimage_subset_range _ ?_ U this _ (fun i ↦ (𝓤.map i).base) + (fun i ↦ (𝓤.map i).continuous) ?_ ?_ + · fun_prop + · rintro i x ⟨⟨x₁, hx₁⟩, ⟨x₂, hx₂⟩⟩ + obtain ⟨x₁', hx₁'⟩ := + Scheme.IsJointlySurjectivePreserving.exists_preimage_fst_triplet_of_prop + (P := @IsOpenImmersion) inferInstance _ _ hx₁.symm + obtain ⟨x₂', hx₂'⟩ := + Scheme.IsJointlySurjectivePreserving.exists_preimage_fst_triplet_of_prop + (P := @IsOpenImmersion) inferInstance _ _ hx₂.symm + obtain ⟨z, hz⟩ := + Scheme.IsJointlySurjectivePreserving.exists_preimage_fst_triplet_of_prop + (P := @IsOpenImmersion) inferInstance _ _ (hx₁'.trans hx₂'.symm) + refine ⟨(pullbackFstFstIso _ _ _ _ _ _ (𝒰.map i.1) ?_ ?_).hom.base z, ?_⟩ + · simp [pullback.condition] + · simp [pullback.condition] + · dsimp only + rw [← hx₁', ← hz, ← Scheme.comp_base_apply] + erw [← Scheme.comp_base_apply] + congr 4 + apply pullback.hom_ext <;> simp [𝓤, ← pullback.condition, ← pullback.condition_assoc] + · intro i + have := H (S.affineOpenCover.obj i.1) (((𝒰.pullbackCover f).obj i.1).affineOpenCover.obj i.2.1) + (((𝒰.pullbackCover g).obj i.1).affineOpenCover.obj i.2.2) + ((𝒱 i.1).map i.2.1 ≫ 𝒰.pullbackHom f i.1) + ((𝒲 i.1).map i.2.2 ≫ 𝒰.pullbackHom g i.1) + ((𝒱 i.1).map i.2.1 ≫ (𝒰.pullbackCover f).map i.1) + ((𝒲 i.1).map i.2.2 ≫ (𝒰.pullbackCover g).map i.1) + (𝒰.map i.1) (by simp [pullback.condition]) (by simp [pullback.condition]) + inferInstance inferInstance inferInstance + convert this using 6 + apply pullback.hom_ext <;> + simp [𝓤, ← pullback.condition, ← pullback.condition_assoc, Scheme.Cover.pullbackHom] + +end SurjectiveOnStalks + +end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Scheme.lean b/Mathlib/AlgebraicGeometry/Scheme.lean index f6caa6dd8c1f7..6decb4e3cb660 100644 --- a/Mathlib/AlgebraicGeometry/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/Scheme.lean @@ -214,6 +214,14 @@ def forgetToTop : Scheme ⥤ TopCat := noncomputable def homeoOfIso {X Y : Scheme.{u}} (e : X ≅ Y) : X ≃ₜ Y := TopCat.homeoOfIso (forgetToTop.mapIso e) +@[simp] +lemma homeoOfIso_symm {X Y : Scheme} (e : X ≅ Y) : + (homeoOfIso e).symm = homeoOfIso e.symm := rfl + +@[simp] +lemma homeoOfIso_apply {X Y : Scheme} (e : X ≅ Y) (x : X) : + homeoOfIso e x = e.hom.base x := rfl + alias _root_.CategoryTheory.Iso.schemeIsoToHomeo := homeoOfIso /-- An isomorphism of schemes induces a homeomorphism of the underlying topological spaces. -/ @@ -221,6 +229,10 @@ noncomputable def Hom.homeomorph {X Y : Scheme.{u}} (f : X.Hom Y) [IsIso (C := S X ≃ₜ Y := (asIso f).schemeIsoToHomeo +@[simp] +lemma Hom.homeomorph_apply {X Y : Scheme.{u}} (f : X.Hom Y) [IsIso (C := Scheme) f] (x) : + f.homeomorph x = f.base x := rfl + -- Porting note: Lean seems not able to find this coercion any more instance hasCoeToTopCat : CoeOut Scheme TopCat where coe X := X.carrier diff --git a/Mathlib/Topology/LocalAtTarget.lean b/Mathlib/Topology/LocalAtTarget.lean index 9d801889f8176..9ab72631191fb 100644 --- a/Mathlib/Topology/LocalAtTarget.lean +++ b/Mathlib/Topology/LocalAtTarget.lean @@ -185,6 +185,47 @@ theorem isEmbedding_iff_of_iSup_eq_top (h : Continuous f) : convert congr_arg SetLike.coe hU simp +omit hU in +/-- +Given a continuous map `f : X → Y` between topological spaces. +Suppose we have an open cover `V i` of the range of `f`, and an open cover `U i` of `X` that is +coarser than the pullback of `V` under `f`. +To check that `f` is an embedding it suffices to check that `U i → Y` is an embedding for all `i`. +-/ +theorem isEmbedding_of_iSup_eq_top_of_preimage_subset_range + {X Y} [TopologicalSpace X] [TopologicalSpace Y] + (f : X → Y) (h : Continuous f) {ι : Type*} + (U : ι → Opens Y) (hU : Set.range f ⊆ (iSup U : _)) + (V : ι → Type*) [∀ i, TopologicalSpace (V i)] + (iV : ∀ i, V i → X) (hiV : ∀ i, Continuous (iV i)) (hV : ∀ i, f ⁻¹' U i ⊆ Set.range (iV i)) + (hV' : ∀ i, IsEmbedding (f ∘ iV i)) : IsEmbedding f := by + wlog hU' : iSup U = ⊤ + · let f₀ : X → Set.range f := fun x ↦ ⟨f x, ⟨x, rfl⟩⟩ + suffices IsEmbedding f₀ from IsEmbedding.subtypeVal.comp this + have hU'' : (⨆ i, (U i).comap ⟨Subtype.val, continuous_subtype_val⟩ : + Opens (Set.range f)) = ⊤ := by + rw [← top_le_iff] + simpa [Set.range_subset_iff, SetLike.le_def] using hU + refine this _ ?_ _ ?_ V iV hiV ?_ ?_ hU'' + · fun_prop + · rw [hU'']; simp + · exact hV + · exact fun i ↦ IsEmbedding.of_comp (by fun_prop) continuous_subtype_val (hV' i) + rw [isEmbedding_iff_of_iSup_eq_top hU' h] + intro i + let f' := (Subtype.val ∘ (f ⁻¹' U i).restrictPreimage (iV i)) + have : IsEmbedding f' := + IsEmbedding.subtypeVal.comp ((IsEmbedding.of_comp (hiV i) h (hV' _)).restrictPreimage _) + have hf' : Set.range f' = f ⁻¹' U i := by + simpa [f', Set.range_comp, Set.range_restrictPreimage] using hV i + let e := (Homeomorph.ofIsEmbedding _ this).trans (Homeomorph.setCongr hf') + refine IsEmbedding.of_comp (by fun_prop) continuous_subtype_val ?_ + convert ((hV' i).comp IsEmbedding.subtypeVal).comp e.symm.isEmbedding + ext x + obtain ⟨x, rfl⟩ := e.surjective x + simp + rfl + @[deprecated (since := "2024-10-26")] alias embedding_iff_embedding_of_iSup_eq_top := isEmbedding_iff_of_iSup_eq_top From 6343082c7dd00a42b18e306d26a4fa513cd044d3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sat, 23 Nov 2024 14:58:16 +0000 Subject: [PATCH 247/829] feat: Sedrakyan's lemma (#19311) This is a specialization of the Cauchy-Schwarz inequality which is often useful in math olympiad problems. In order to prove it in general ordered semifields, we have to show slightly more general forms of the AM-GM and Cauchy-Schwarz inequalities, which indirectly work with square roots. Co-authored-by: Alex Brodbelt <64128135+AlexBrodbelt@users.noreply.github.com> Co-authored-by: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> --- .../Order/BigOperators/Ring/Finset.lean | 102 +++++++++++------- .../Algebra/Order/Ring/Unbundled/Basic.lean | 12 ++- 2 files changed, 75 insertions(+), 39 deletions(-) diff --git a/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean b/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean index b51accd2d82ad..1c522e8948d04 100644 --- a/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean @@ -14,9 +14,12 @@ import Mathlib.Tactic.Ring This file contains the results concerning the interaction of finset big operators with ordered rings. + +In particular, this file contains the standard form of the Cauchy-Schwarz inequality, as well as +some of its immediate consequences. -/ -variable {ι R : Type*} +variable {ι R S : Type*} namespace Finset @@ -120,41 +123,10 @@ lemma prod_add_prod_le {i : ι} {f g h : ι → R} (hi : i ∈ s) (h2i : g i + h end OrderedCommSemiring -section LinearOrderedCommSemiring -variable [LinearOrderedCommSemiring R] [ExistsAddOfLE R] - -/-- **Cauchy-Schwarz inequality** for finsets. -/ -lemma sum_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 - nontriviality R - obtain h' | h' := (sum_nonneg fun _ _ ↦ sq_nonneg <| g _).eq_or_lt - · have h'' : ∀ i ∈ s, g i = 0 := fun i hi ↦ by - simpa using (sum_eq_zero_iff_of_nonneg fun i _ ↦ sq_nonneg (g i)).1 h'.symm i hi - rw [← h', sum_congr rfl (show ∀ i ∈ s, f i * g i = 0 from fun i hi ↦ by simp [h'' i hi])] - simp - refine le_of_mul_le_mul_of_pos_left - (le_of_add_le_add_left (a := (∑ i ∈ s, g i ^ 2) * (∑ j ∈ s, f j * g j) ^ 2) ?_) h' - calc - _ = ∑ i ∈ s, 2 * (f i * ∑ j ∈ s, g j ^ 2) * (g i * ∑ j ∈ s, f j * g j) := by - simp_rw [mul_assoc (2 : R), mul_mul_mul_comm, ← mul_sum, ← sum_mul]; ring - _ ≤ ∑ i ∈ s, ((f i * ∑ j ∈ s, g j ^ 2) ^ 2 + (g i * ∑ j ∈ s, f j * g j) ^ 2) := - sum_le_sum fun i _ ↦ two_mul_le_add_sq (f i * ∑ j ∈ s, g j ^ 2) (g i * ∑ j ∈ s, f j * g j) - _ = _ := by simp_rw [sum_add_distrib, mul_pow, ← sum_mul]; ring - -theorem sum_mul_self_eq_zero_iff (s : Finset ι) (f : ι → R) : - ∑ i ∈ s, f i * f i = 0 ↔ ∀ i ∈ s, f i = 0 := by - induction s using Finset.cons_induction with - | empty => simp - | cons i s his ih => - simp only [Finset.sum_cons, Finset.mem_cons, forall_eq_or_imp] - refine ⟨fun hc => ?_, fun h => by simpa [h.1] using ih.mpr h.2⟩ - have hi : f i * f i ≤ 0 := by - rw [← hc, le_add_iff_nonneg_right] - exact Finset.sum_nonneg fun i _ ↦ mul_self_nonneg (f i) - have h : f i * f i = 0 := (eq_of_le_of_le (mul_self_nonneg (f i)) hi).symm - exact ⟨zero_eq_mul_self.mp h.symm, ih.mp (by rw [← hc, h, zero_add])⟩ - -end LinearOrderedCommSemiring +theorem sum_mul_self_eq_zero_iff [LinearOrderedSemiring R] [ExistsAddOfLE R] (s : Finset ι) + (f : ι → R) : ∑ i ∈ s, f i * f i = 0 ↔ ∀ i ∈ s, f i = 0 := by + rw [sum_eq_zero_iff_of_nonneg fun _ _ ↦ mul_self_nonneg _] + simp lemma abs_prod [LinearOrderedCommRing R] (s : Finset ι) (f : ι → R) : |∏ x ∈ s, f x| = ∏ x ∈ s, |f x| := @@ -184,11 +156,63 @@ lemma prod_add_prod_le' (hi : i ∈ s) (h2i : g i + h i ≤ f i) (hgf : ∀ j assumption end CanonicallyOrderedCommSemiring + +/-! ### Named inequalities -/ + +/-- **Cauchy-Schwarz inequality** for finsets. + +This is written in terms of sequences `f`, `g`, and `r`, where `r` is a stand-in for +`√(f i * g i)`. See `sum_mul_sq_le_sq_mul_sq` for the more usual form in terms of squared +sequences. -/ +lemma sum_sq_le_sum_mul_sum_of_sq_eq_mul [LinearOrderedCommSemiring R] [ExistsAddOfLE R] + (s : Finset ι) {r f g : ι → R} (hf : ∀ i ∈ s, 0 ≤ f i) (hg : ∀ i ∈ s, 0 ≤ g i) + (ht : ∀ i ∈ s, r i ^ 2 = f i * g i) : (∑ i ∈ s, r i) ^ 2 ≤ (∑ i ∈ s, f i) * ∑ i ∈ s, g i := by + obtain h | h := (sum_nonneg hg).eq_or_gt + · have ht' : ∑ i ∈ s, r i = 0 := sum_eq_zero fun i hi ↦ by + simpa [(sum_eq_zero_iff_of_nonneg hg).1 h i hi] using ht i hi + rw [h, ht'] + simp + · refine le_of_mul_le_mul_of_pos_left + (le_of_add_le_add_left (a := (∑ i ∈ s, g i) * (∑ i ∈ s, r i) ^ 2) ?_) h + calc + _ = ∑ i ∈ s, 2 * r i * (∑ j ∈ s, g j) * (∑ j ∈ s, r j) := by + simp_rw [mul_assoc, ← mul_sum, ← sum_mul]; ring + _ ≤ ∑ i ∈ s, (f i * (∑ j ∈ s, g j) ^ 2 + g i * (∑ j ∈ s, r j) ^ 2) := by + gcongr with i hi + have ht : (r i * (∑ j ∈ s, g j) * (∑ j ∈ s, r j)) ^ 2 = + (f i * (∑ j ∈ s, g j) ^ 2) * (g i * (∑ j ∈ s, r j) ^ 2) := by + conv_rhs => rw [mul_mul_mul_comm, ← ht i hi] + ring + refine le_of_eq_of_le ?_ (two_mul_le_add_of_sq_eq_mul + (mul_nonneg (hf i hi) (sq_nonneg _)) (mul_nonneg (hg i hi) (sq_nonneg _)) ht) + repeat rw [mul_assoc] + _ = _ := by simp_rw [sum_add_distrib, ← sum_mul]; ring + +/-- **Cauchy-Schwarz inequality** for finsets, squared version. -/ +lemma sum_mul_sq_le_sq_mul_sq [LinearOrderedCommSemiring R] [ExistsAddOfLE R] (s : Finset ι) + (f g : ι → R) : (∑ i ∈ s, f i * g i) ^ 2 ≤ (∑ i ∈ s, f i ^ 2) * ∑ i ∈ s, g i ^ 2 := + sum_sq_le_sum_mul_sum_of_sq_eq_mul s + (fun _ _ ↦ sq_nonneg _) (fun _ _ ↦ sq_nonneg _) (fun _ _ ↦ mul_pow ..) + +/-- **Sedrakyan's lemma**, aka **Titu's lemma** or **Engel's form**. + +This is a specialization of the Cauchy-Schwarz inequality with the sequences `f n / √(g n)` and +`√(g n)`, though here it is proven without relying on square roots. -/ +theorem sq_sum_div_le_sum_sq_div [LinearOrderedSemifield R] [ExistsAddOfLE R] (s : Finset ι) + (f : ι → R) {g : ι → R} (hg : ∀ i ∈ s, 0 < g i) : + (∑ i ∈ s, f i) ^ 2 / ∑ i ∈ s, g i ≤ ∑ i ∈ s, f i ^ 2 / g i := by + have hg' : ∀ i ∈ s, 0 ≤ g i := fun i hi ↦ (hg i hi).le + have H : ∀ i ∈ s, 0 ≤ f i ^ 2 / g i := fun i hi ↦ div_nonneg (sq_nonneg _) (hg' i hi) + refine div_le_of_le_mul₀ (sum_nonneg hg') (sum_nonneg H) + (sum_sq_le_sum_mul_sum_of_sq_eq_mul _ H hg' fun i hi ↦ ?_) + rw [div_mul_cancel₀] + exact (hg i hi).ne' + end Finset -section AbsoluteValue +/-! ### Absolute values -/ -variable {S : Type*} +section AbsoluteValue lemma AbsoluteValue.sum_le [Semiring R] [OrderedSemiring S] (abv : AbsoluteValue R S) (s : Finset ι) (f : ι → R) : abv (∑ i ∈ s, f i) ≤ ∑ i ∈ s, abv (f i) := @@ -212,6 +236,8 @@ lemma IsAbsoluteValue.map_prod [CommSemiring R] [Nontrivial R] [LinearOrderedCom end AbsoluteValue +/-! ### Positivity extension -/ + namespace Mathlib.Meta.Positivity open Qq Lean Meta Finset diff --git a/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean index 06cc0315f08b0..99bd1f2cfd31c 100644 --- a/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean @@ -764,7 +764,7 @@ alias two_mul_le_add_pow_two := two_mul_le_add_sq /-- Binary, squared, and division-free **arithmetic mean-geometric mean inequality** (aka AM-GM inequality) for linearly ordered commutative semirings. -/ lemma four_mul_le_sq_add [ExistsAddOfLE α] [MulPosStrictMono α] - [ContravariantClass α α (· + ·) (· ≤ ·)] [CovariantClass α α (· + ·) (· ≤ ·)] + [AddLeftReflectLE α] [AddLeftMono α] (a b : α) : 4 * a * b ≤ (a + b) ^ 2 := by calc 4 * a * b _ = 2 * a * b + 2 * a * b := by rw [mul_assoc, two_add_two_eq_four.symm, add_mul, mul_assoc] @@ -774,6 +774,16 @@ lemma four_mul_le_sq_add [ExistsAddOfLE α] [MulPosStrictMono α] alias four_mul_le_pow_two_add := four_mul_le_sq_add +/-- Binary and division-free **arithmetic mean-geometric mean inequality** +(aka AM-GM inequality) for linearly ordered commutative semirings. -/ +lemma two_mul_le_add_of_sq_eq_mul [ExistsAddOfLE α] [MulPosStrictMono α] [PosMulStrictMono α] + [AddLeftReflectLE α] [AddLeftMono α] {a b r : α} + (ha : 0 ≤ a) (hb : 0 ≤ b) (ht : r ^ 2 = a * b) : 2 * r ≤ a + b := by + apply nonneg_le_nonneg_of_sq_le_sq (Left.add_nonneg ha hb) + conv_rhs => rw [← pow_two] + convert four_mul_le_sq_add a b using 1 + rw [mul_mul_mul_comm, two_mul, two_add_two_eq_four, ← pow_two, ht, mul_assoc] + end LinearOrderedCommSemiring section LinearOrderedRing From cc6b7ff0c02212e54b56437933156ee1cdf452fe Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 23 Nov 2024 15:57:51 +0000 Subject: [PATCH 248/829] =?UTF-8?q?feat:=20`s=E2=81=BB=C2=B9.encard=20=3D?= =?UTF-8?q?=20s.encard`=20(#19400)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From Kneser (LeanCamCombi) --- Mathlib/Algebra/Group/Pointwise/Set/Card.lean | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean index 80a073b792385..06a3718af24d0 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean @@ -3,6 +3,7 @@ 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.Data.Set.Card import Mathlib.Data.Set.Pointwise.Finite import Mathlib.SetTheory.Cardinal.Finite @@ -49,6 +50,12 @@ lemma natCard_inv (s : Set G) : Nat.card ↥(s⁻¹) = Nat.card s := by @[to_additive (attr := deprecated (since := "2024-09-30"))] alias card_inv := natCard_inv +@[to_additive (attr := simp)] +lemma encard_inv (s : Set G) : s⁻¹.encard = s.encard := by simp [encard, PartENat.card] + +@[to_additive (attr := simp)] +lemma ncard_inv (s : Set G) : s⁻¹.ncard = s.ncard := by simp [ncard] + end InvolutiveInv section DivInvMonoid @@ -82,5 +89,12 @@ lemma natCard_smul_set (a : G) (s : Set α) : Nat.card ↥(a • s) = Nat.card s @[to_additive (attr := deprecated (since := "2024-09-30"))] alias card_smul_set := Cardinal.mk_smul_set +@[to_additive (attr := simp)] +lemma encard_smul_set (a : G) (s : Set α) : (a • s).encard = s.encard := by + simp [encard, PartENat.card] + +@[to_additive (attr := simp)] +lemma ncard_smul_set (a : G) (s : Set α) : (a • s).ncard = s.ncard := by simp [ncard] + end Group end Set From 571c2829083a7dcf8934dff094f89869b3f6244f Mon Sep 17 00:00:00 2001 From: damiano Date: Sat, 23 Nov 2024 19:35:47 +0000 Subject: [PATCH 249/829] chore: remove unused `Category` assumption (#19416) The `Category` assumption is unused, as `Discrete` subsumes it. --- Mathlib/CategoryTheory/Groupoid/Discrete.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Groupoid/Discrete.lean b/Mathlib/CategoryTheory/Groupoid/Discrete.lean index ffface4377069..3ce9a350831bf 100644 --- a/Mathlib/CategoryTheory/Groupoid/Discrete.lean +++ b/Mathlib/CategoryTheory/Groupoid/Discrete.lean @@ -12,7 +12,7 @@ import Mathlib.CategoryTheory.DiscreteCategory namespace CategoryTheory -variable {C : Type*} [Category C] +variable {C : Type*} instance : Groupoid (Discrete C) := { inv := fun h ↦ ⟨⟨h.1.1.symm⟩⟩ } From fee45b61fbab23192de1eb6f8f536e53722254f1 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Sat, 23 Nov 2024 19:55:50 +0000 Subject: [PATCH 250/829] feat(AlgebraicGeometry): the domain of definition of a rational map (#18803) --- Mathlib/AlgebraicGeometry/AffineSpace.lean | 2 +- Mathlib/AlgebraicGeometry/Limits.lean | 6 +- .../Morphisms/Separated.lean | 10 + Mathlib/AlgebraicGeometry/Over.lean | 4 + Mathlib/AlgebraicGeometry/RationalMap.lean | 220 +++++++++++++++++- Mathlib/AlgebraicGeometry/Restrict.lean | 21 +- Mathlib/CategoryTheory/Comma/OverClass.lean | 20 +- 7 files changed, 273 insertions(+), 10 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/AffineSpace.lean b/Mathlib/AlgebraicGeometry/AffineSpace.lean index 6374a6564ad9f..8e43ca277f77f 100644 --- a/Mathlib/AlgebraicGeometry/AffineSpace.lean +++ b/Mathlib/AlgebraicGeometry/AffineSpace.lean @@ -332,7 +332,7 @@ def mapSpecMap {R S : CommRingCat.{max u v}} (φ : R ⟶ S) : def reindex {n m : Type v} (i : m → n) (S : Scheme.{max u v}) : 𝔸(n; S) ⟶ 𝔸(m; S) := homOfVector (𝔸(n; S) ↘ S) (coord S ∘ i) -@[reassoc (attr := simp)] +@[simp, reassoc] lemma reindex_over {n m : Type v} (i : m → n) (S : Scheme.{max u v}) : reindex i S ≫ 𝔸(m; S) ↘ S = 𝔸(n; S) ↘ S := pullback.lift_fst _ _ _ diff --git a/Mathlib/AlgebraicGeometry/Limits.lean b/Mathlib/AlgebraicGeometry/Limits.lean index 6c634b3aadaac..b4f7a0106ddca 100644 --- a/Mathlib/AlgebraicGeometry/Limits.lean +++ b/Mathlib/AlgebraicGeometry/Limits.lean @@ -48,7 +48,11 @@ instance : HasFiniteLimits Scheme := hasFiniteLimits_of_hasTerminal_and_pullbacks instance (X : Scheme.{u}) : X.Over (⊤_ _) := ⟨terminal.from _⟩ -instance {X Y : Scheme.{u}} (f : X ⟶ Y) : f.IsOver (⊤_ _) := ⟨terminal.hom_ext _ _⟩ +instance {X Y : Scheme.{u}} [X.Over (⊤_ Scheme)] [Y.Over (⊤_ Scheme)] (f : X ⟶ Y) : + @Scheme.Hom.IsOver _ _ f (⊤_ Scheme) ‹_› ‹_› := ⟨Subsingleton.elim _ _⟩ + +instance {X : Scheme} : Subsingleton (X.Over (⊤_ Scheme)) := + ⟨fun ⟨a⟩ ⟨b⟩ ↦ by simp [Subsingleton.elim a b]⟩ section Initial diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean b/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean index a97b12f9fcc9c..c03365e68f13f 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean @@ -240,6 +240,16 @@ lemma ext_of_isDominant_of_isSeparated [IsReduced X] {f g : X ⟶ Y} rw [← cancel_epi (equalizer.ι f' g').left] exact congr($(equalizer.condition f' g').left) +variable (S) in +/-- +Suppose `X` is a reduced `S`-scheme and `Y` is a separated `S`-scheme. +For any `S`-morphisms `f g : X ⟶ Y`, `f = g` if `ι ≫ f = ι ≫ g` for some dominant `ι`. +-/ +lemma ext_of_isDominant_of_isSeparated' [X.Over S] [Y.Over S] [IsReduced X] [IsSeparated (Y ↘ S)] + {f g : X ⟶ Y} [f.IsOver S] [g.IsOver S] {W} (ι : W ⟶ X) [IsDominant ι] + (hU : ι ≫ f = ι ≫ g) : f = g := + ext_of_isDominant_of_isSeparated (Y ↘ S) (by simp) ι hU + namespace Scheme /-- A scheme `X` is separated if it is separated over `⊤_ Scheme`. -/ diff --git a/Mathlib/AlgebraicGeometry/Over.lean b/Mathlib/AlgebraicGeometry/Over.lean index c52efcbb93841..b0f3bcaef6d0e 100644 --- a/Mathlib/AlgebraicGeometry/Over.lean +++ b/Mathlib/AlgebraicGeometry/Over.lean @@ -41,6 +41,10 @@ abbrev CanonicallyOver := CanonicallyOverClass X S `f.IsOver S` is the typeclass asserting `f` commutes with the structure morphisms. -/ abbrev Hom.IsOver (f : X.Hom Y) (S : Scheme.{u}) [X.Over S] [Y.Over S] := HomIsOver f S +@[simp] +lemma Hom.isOver_iff [X.Over S] [Y.Over S] {f : X ⟶ Y} : f.IsOver S ↔ f ≫ Y ↘ S = X ↘ S := + ⟨fun H ↦ H.1, fun h ↦ ⟨h⟩⟩ + /-! Also note the existence of `CategoryTheory.IsOverTower X Y S`. -/ end AlgebraicGeometry.Scheme diff --git a/Mathlib/AlgebraicGeometry/RationalMap.lean b/Mathlib/AlgebraicGeometry/RationalMap.lean index aca42607f4a77..e9894ddfbcd61 100644 --- a/Mathlib/AlgebraicGeometry/RationalMap.lean +++ b/Mathlib/AlgebraicGeometry/RationalMap.lean @@ -5,6 +5,7 @@ Authors: Andrew Yang -/ import Mathlib.AlgebraicGeometry.SpreadingOut import Mathlib.AlgebraicGeometry.FunctionField +import Mathlib.AlgebraicGeometry.Morphisms.Separated /-! # Rational maps between schemes @@ -17,10 +18,12 @@ import Mathlib.AlgebraicGeometry.FunctionField Two partial maps are equivalent if they are equal on a dense open subscheme. * `AlgebraicGeometry.Scheme.RationalMap`: A rational map from `X` to `Y` (`X ⤏ Y`) is an equivalence class of partial maps. -* `AlgebraicGeometry.Scheme.RationalMap.equivFunctionField`: +* `AlgebraicGeometry.Scheme.RationalMap.equivFunctionFieldOver`: Given `S`-schemes `X` and `Y` such that `Y` is locally of finite type and `X` is integral, `S`-morphisms `Spec K(X) ⟶ Y` correspond bijectively to `S`-rational maps from `X` to `Y`. - +* `AlgebraicGeometry.Scheme.RationalMap.toPartialMap`: + If `X` is integral and `Y` is separated, then any `f : X ⤏ Y` can be realized as a partial + map on `f.domain`, the domain of definition of `f`. -/ universe u @@ -44,6 +47,11 @@ structure PartialMap (X Y : Scheme.{u}) where /-- The underlying morphism of a partial map. -/ hom : ↑domain ⟶ Y +variable (S) in +/-- A partial map is a `S`-map if the underlying morphism is. -/ +abbrev PartialMap.IsOver [X.Over S] [Y.Over S] (f : X.PartialMap Y) := + f.hom.IsOver S + namespace PartialMap lemma ext_iff (f g : X.PartialMap Y) : @@ -64,7 +72,7 @@ lemma ext (f g : X.PartialMap Y) (e : f.domain = g.domain) exact ⟨e, H⟩ /-- The restriction of a partial map to a smaller domain. -/ -@[simps hom, simps (config := .lemmasOnly) domain] +@[simps hom domain] noncomputable def restrict (f : X.PartialMap Y) (U : X.Opens) (hU : Dense (U : Set X)) (hU' : U ≤ f.domain) : X.PartialMap Y where @@ -93,6 +101,10 @@ lemma restrict_restrict_hom (f : X.PartialMap Y) ((f.restrict U hU hU').restrict V hV hV').hom = (f.restrict V hV (hV'.trans hU')).hom := by simp +instance [X.Over S] [Y.Over S] (f : X.PartialMap Y) [f.IsOver S] + (U : X.Opens) (hU : Dense (U : Set X)) (hU' : U ≤ f.domain) : + (f.restrict U hU hU').IsOver S where + /-- The composition of a partial map and a morphism on the right. -/ @[simps] def compHom (f : X.PartialMap Y) (g : Y ⟶ Z) : X.PartialMap Z where @@ -100,11 +112,24 @@ def compHom (f : X.PartialMap Y) (g : Y ⟶ Z) : X.PartialMap Z where dense_domain := f.dense_domain hom := f.hom ≫ g +instance [X.Over S] [Y.Over S] [Z.Over S] (f : X.PartialMap Y) (g : Y ⟶ Z) + [f.IsOver S] [g.IsOver S] : (f.compHom g).IsOver S where + /-- A scheme morphism as a partial map. -/ @[simps] def _root_.AlgebraicGeometry.Scheme.Hom.toPartialMap (f : X.Hom Y) : X.PartialMap Y := ⟨⊤, dense_univ, X.topIso.hom ≫ f⟩ +instance [X.Over S] [Y.Over S] (f : X ⟶ Y) [f.IsOver S] : f.toPartialMap.IsOver S where + +lemma isOver_iff [X.Over S] [Y.Over S] {f : X.PartialMap Y} : + f.IsOver S ↔ (f.compHom (Y ↘ S)).hom = f.domain.ι ≫ X ↘ S := by + simp + +lemma isOver_iff_eq_restrict [X.Over S] [Y.Over S] {f : X.PartialMap Y} : + f.IsOver S ↔ f.compHom (Y ↘ S) = (X ↘ S).toPartialMap.restrict _ f.dense_domain (by simp) := by + simp [isOver_iff, PartialMap.ext_iff] + /-- If `x` is in the domain of a partial map `f`, then `f` restricts to a map from `Spec 𝒪_x`. -/ noncomputable def fromSpecStalkOfMem (f : X.PartialMap Y) {x} (hx : x ∈ f.domain) : @@ -230,6 +255,58 @@ lemma equiv_of_fromSpecStalkOfMem_eq [IrreducibleSpace X] simpa only [fromSpecStalkOfMem, restrict_domain, Opens.fromSpecStalkOfMem, Spec.map_inv, restrict_hom, Category.assoc, IsIso.eq_inv_comp, IsIso.hom_inv_id_assoc] using H +instance (U : X.Opens) [IsReduced X] : IsReduced U := isReduced_of_isOpenImmersion U.ι + +lemma Opens.isDominant_ι {U : X.Opens} (hU : Dense (X := X) U) : IsDominant U.ι := + ⟨by simpa [DenseRange] using hU⟩ + +lemma Opens.isDominant_homOfLE {U V : X.Opens} (hU : Dense (X := X) U) (hU' : U ≤ V) : + IsDominant (X.homOfLE hU') := + have : IsDominant (X.homOfLE hU' ≫ Opens.ι _) := by simpa using Opens.isDominant_ι hU + IsDominant.of_comp_of_isOpenImmersion (g := Opens.ι _) _ + +/-- Two partial maps from reduced schemes to separated schemes are equivalent if and only if +they are equal on **any** open dense subset. -/ +lemma equiv_iff_of_isSeparated_of_le [X.Over S] [Y.Over S] [IsReduced X] + [IsSeparated (Y ↘ S)] {f g : X.PartialMap Y} [f.IsOver S] [g.IsOver S] + {W : X.Opens} (hW : Dense (X := X) W) (hWl : W ≤ f.domain) (hWr : W ≤ g.domain) : f.equiv g ↔ + (f.restrict W hW hWl).hom = (g.restrict W hW hWr).hom := by + refine ⟨fun ⟨V, hV, hVl, hVr, e⟩ ↦ ?_, fun e ↦ ⟨_, _, _, _, e⟩⟩ + have : IsDominant (X.homOfLE (inf_le_left : W ⊓ V ≤ W)) := + Opens.isDominant_homOfLE (hW.inter_of_isOpen_left hV W.2) _ + apply ext_of_isDominant_of_isSeparated' S (X.homOfLE (inf_le_left : W ⊓ V ≤ W)) + simpa using congr(X.homOfLE (inf_le_right : W ⊓ V ≤ V) ≫ $e) + +/-- Two partial maps from reduced schemes to separated schemes are equivalent if and only if +they are equal on the intersection of the domains. -/ +lemma equiv_iff_of_isSeparated [X.Over S] [Y.Over S] [IsReduced X] + [IsSeparated (Y ↘ S)] {f g : X.PartialMap Y} + [f.IsOver S] [g.IsOver S] : f.equiv g ↔ + (f.restrict _ (f.2.inter_of_isOpen_left g.2 f.domain.2) inf_le_left).hom = + (g.restrict _ (f.2.inter_of_isOpen_left g.2 f.domain.2) inf_le_right).hom := + equiv_iff_of_isSeparated_of_le (S := S) _ _ _ + +/-- Two partial maps from reduced schemes to separated schemes with the same domain are equivalent +if and only if they are equal. -/ +lemma equiv_iff_of_domain_eq_of_isSeparated [X.Over S] [Y.Over S] [IsReduced X] + [IsSeparated (Y ↘ S)] {f g : X.PartialMap Y} (hfg : f.domain = g.domain) + [f.IsOver S] [g.IsOver S] : f.equiv g ↔ f = g := by + rw [equiv_iff_of_isSeparated_of_le (S := S) f.dense_domain le_rfl hfg.le] + obtain ⟨Uf, _, f⟩ := f + obtain ⟨Ug, _, g⟩ := g + obtain rfl : Uf = Ug := hfg + simp + +/-- A partial map from a reduced scheme to a separated scheme is equivalent to a morphism +if and only if it is equal to the restriction of the morphism. -/ +lemma equiv_toPartialMap_iff_of_isSeparated [X.Over S] [Y.Over S] [IsReduced X] + [IsSeparated (Y ↘ S)] {f : X.PartialMap Y} {g : X ⟶ Y} + [f.IsOver S] [g.IsOver S] : f.equiv g.toPartialMap ↔ + f.hom = f.domain.ι ≫ g := by + rw [equiv_iff_of_isSeparated (S := S), ← cancel_epi (X.isoOfEq (inf_top_eq f.domain)).hom] + simp + rfl + end PartialMap /-- A rational map from `X` to `Y` (`X ⤏ Y`) is an equivalence class of partial maps, @@ -246,6 +323,11 @@ def PartialMap.toRationalMap (f : X.PartialMap Y) : X ⤏ Y := Quotient.mk _ f /-- A scheme morphism as a rational map. -/ abbrev Hom.toRationalMap (f : X.Hom Y) : X ⤏ Y := f.toPartialMap.toRationalMap +variable (S) in +/-- A rational map is a `S`-map if some partial map in the equivalence class is a `S`-map. -/ +class RationalMap.IsOver [X.Over S] [Y.Over S] (f : X ⤏ Y) : Prop where + exists_partialMap_over : ∃ g : X.PartialMap Y, g.IsOver S ∧ g.toRationalMap = f + lemma PartialMap.toRationalMap_surjective : Function.Surjective (@toRationalMap X Y) := Quotient.exists_rep @@ -262,6 +344,14 @@ lemma PartialMap.restrict_toRationalMap (f : X.PartialMap Y) (U : X.Opens) (f.restrict U hU hU').toRationalMap = f.toRationalMap := toRationalMap_eq_iff.mpr (f.restrict_equiv U hU hU') +instance [X.Over S] [Y.Over S] (f : X.PartialMap Y) [f.IsOver S] : f.toRationalMap.IsOver S := + ⟨f, ‹_›, rfl⟩ + +variable (S) in +lemma RationalMap.exists_partialMap_over [X.Over S] [Y.Over S] (f : X ⤏ Y) [f.IsOver S] : + ∃ g : X.PartialMap Y, g.IsOver S ∧ g.toRationalMap = f := + IsOver.exists_partialMap_over + /-- The composition of a rational map and a morphism on the right. -/ def RationalMap.compHom (f : X ⤏ Y) (g : Y ⟶ Z) : X ⤏ Z := by refine Quotient.map (PartialMap.compHom · g) ?_ f @@ -275,6 +365,42 @@ def RationalMap.compHom (f : X ⤏ Y) (g : Y ⟶ Z) : X ⤏ Z := by lemma RationalMap.compHom_toRationalMap (f : X.PartialMap Y) (g : Y ⟶ Z) : (f.compHom g).toRationalMap = f.toRationalMap.compHom g := rfl +instance [X.Over S] [Y.Over S] [Z.Over S] (f : X ⤏ Y) (g : Y ⟶ Z) + [f.IsOver S] [g.IsOver S] : (f.compHom g).IsOver S where + exists_partialMap_over := by + obtain ⟨f, hf, rfl⟩ := f.exists_partialMap_over S + exact ⟨f.compHom g, inferInstance, rfl⟩ + +variable (S) in +lemma PartialMap.exists_restrict_isOver [X.Over S] [Y.Over S] (f : X.PartialMap Y) + [f.toRationalMap.IsOver S] : ∃ U hU hU', (f.restrict U hU hU').IsOver S := by + obtain ⟨f', hf₁, hf₂⟩ := RationalMap.IsOver.exists_partialMap_over (S := S) (f := f.toRationalMap) + obtain ⟨U, hU, hUl, hUr, e⟩ := PartialMap.toRationalMap_eq_iff.mp hf₂ + exact ⟨U, hU, hUr, by rw [IsOver, ← e]; infer_instance⟩ + +lemma RationalMap.isOver_iff [X.Over S] [Y.Over S] {f : X ⤏ Y} : + f.IsOver S ↔ f.compHom (Y ↘ S) = (X ↘ S).toRationalMap := by + constructor + · intro h + obtain ⟨g, hg, e⟩ := f.exists_partialMap_over S + rw [← e, Hom.toRationalMap, ← compHom_toRationalMap, PartialMap.isOver_iff_eq_restrict.mp hg, + PartialMap.restrict_toRationalMap] + · intro e + obtain ⟨f, rfl⟩ := PartialMap.toRationalMap_surjective f + obtain ⟨U, hU, hUl, hUr, e⟩ := PartialMap.toRationalMap_eq_iff.mp e + exact ⟨⟨f.restrict U hU hUl, by simpa using e, by simp⟩⟩ + +lemma PartialMap.isOver_toRationalMap_iff_of_isSeparated [X.Over S] [Y.Over S] [IsReduced X] + [S.IsSeparated] {f : X.PartialMap Y} : + f.toRationalMap.IsOver S ↔ f.IsOver S := by + refine ⟨fun _ ↦ ?_, fun _ ↦ inferInstance⟩ + obtain ⟨U, hU, hU', H⟩ := f.exists_restrict_isOver (S := S) + rw [isOver_iff] + have : IsDominant (X.homOfLE hU') := Opens.isDominant_homOfLE hU _ + exact ext_of_isDominant (ι := X.homOfLE hU') (by simpa using H.1) + +section functionField + /-- A rational map restricts to a map from `Spec K(X)`. -/ noncomputable def RationalMap.fromFunctionField [IrreducibleSpace X] (f : X ⤏ Y) : @@ -328,6 +454,94 @@ def RationalMap.equivFunctionField [IsIntegral X] [LocallyOfFiniteType sY] : (ofFunctionField sX sY f.1.fromFunctionField _) f (RationalMap.fromFunctionField_ofFunctionField _ _ _ _)) +/-- +Given `S`-schemes `X` and `Y` such that `Y` is locally of finite type and `X` is integral, +`S`-morphisms `Spec K(X) ⟶ Y` correspond bijectively to `S`-rational maps from `X` to `Y`. +-/ +noncomputable +def RationalMap.equivFunctionFieldOver [X.Over S] [Y.Over S] [IsIntegral X] + [LocallyOfFiniteType (Y ↘ S)] : + { f : Spec X.functionField ⟶ Y // f.IsOver S } ≃ { f : X ⤏ Y // f.IsOver S } := + ((Equiv.subtypeEquivProp (by simp only [Hom.isOver_iff]; rfl)).trans + (RationalMap.equivFunctionField (X ↘ S) (Y ↘ S))).trans + (Equiv.subtypeEquivProp (by ext f; rw [RationalMap.isOver_iff])) + +end functionField + +section domain + +/-- The domain of definition of a rational map. -/ +def RationalMap.domain (f : X ⤏ Y) : X.Opens := + sSup { PartialMap.domain g | (g) (_ : g.toRationalMap = f) } + +lemma PartialMap.le_domain_toRationalMap (f : X.PartialMap Y) : + f.domain ≤ f.toRationalMap.domain := + le_sSup ⟨f, rfl, rfl⟩ + +lemma RationalMap.mem_domain {f : X ⤏ Y} {x} : + x ∈ f.domain ↔ ∃ g : X.PartialMap Y, x ∈ g.domain ∧ g.toRationalMap = f := + TopologicalSpace.Opens.mem_sSup.trans (by simp [@and_comm (x ∈ _)]) + +lemma RationalMap.dense_domain (f : X ⤏ Y) : Dense (X := X) f.domain := + f.inductionOn (fun g ↦ g.dense_domain.mono g.le_domain_toRationalMap) + +/-- The open cover of the domain of `f : X ⤏ Y`, +consisting of all the domains of the partial maps in the equivalence class. -/ +noncomputable +def RationalMap.openCoverDomain (f : X ⤏ Y) : f.domain.toScheme.OpenCover where + J := { PartialMap.domain g | (g) (_ : g.toRationalMap = f) } + obj U := U.1.toScheme + map U := X.homOfLE (le_sSup U.2) + f x := ⟨_, (TopologicalSpace.Opens.mem_sSup.mp x.2).choose_spec.1⟩ + covers x := ⟨⟨x.1, (TopologicalSpace.Opens.mem_sSup.mp x.2).choose_spec.2⟩, Subtype.ext (by simp)⟩ + +/-- If `f : X ⤏ Y` is a rational map from a reduced scheme to a separated scheme, +then `f` can be represented as a partial map on its domain of definition. -/ +noncomputable +def RationalMap.toPartialMap [IsReduced X] [Y.IsSeparated] (f : X ⤏ Y) : X.PartialMap Y := by + refine ⟨f.domain, f.dense_domain, f.openCoverDomain.glueMorphisms + (fun x ↦ (X.isoOfEq x.2.choose_spec.2).inv ≫ x.2.choose.hom) ?_⟩ + intros x y + let g (x : f.openCoverDomain.J) := x.2.choose + have hg₁ (x) : (g x).toRationalMap = f := x.2.choose_spec.1 + have hg₂ (x) : (g x).domain = x.1 := x.2.choose_spec.2 + refine (cancel_epi (isPullback_opens_inf_le (le_sSup x.2) (le_sSup y.2)).isoPullback.hom).mp ?_ + simp only [openCoverDomain, IsPullback.isoPullback_hom_fst_assoc, + IsPullback.isoPullback_hom_snd_assoc] + show _ ≫ _ ≫ (g x).hom = _ ≫ _ ≫ (g y).hom + simp_rw [← cancel_epi (X.isoOfEq congr($(hg₂ x) ⊓ $(hg₂ y))).hom, ← Category.assoc] + convert (PartialMap.equiv_iff_of_isSeparated (S := ⊤_ _) (f := g x) (g := g y)).mp ?_ using 1 + · dsimp; congr 1; simp [← cancel_mono (Opens.ι _)] + · dsimp; congr 1; simp [← cancel_mono (Opens.ι _)] + · rw [← PartialMap.toRationalMap_eq_iff, hg₁, hg₁] + +lemma PartialMap.toPartialMap_toRationalMap_restrict [IsReduced X] [Y.IsSeparated] + (f : X.PartialMap Y) : (f.toRationalMap.toPartialMap.restrict _ f.dense_domain + f.le_domain_toRationalMap).hom = f.hom := by + dsimp [RationalMap.toPartialMap] + refine (f.toRationalMap.openCoverDomain.ι_glueMorphisms _ _ ⟨_, f, rfl, rfl⟩).trans ?_ + generalize_proofs _ _ H _ + have : H.choose = f := (equiv_iff_of_domain_eq_of_isSeparated (S := ⊤_ _) H.choose_spec.2).mp + (toRationalMap_eq_iff.mp H.choose_spec.1) + exact ((ext_iff _ _).mp this.symm).choose_spec.symm + +@[simp] +lemma RationalMap.toRationalMap_toPartialMap [IsReduced X] [Y.IsSeparated] + (f : X ⤏ Y) : f.toPartialMap.toRationalMap = f := by + obtain ⟨f, rfl⟩ := PartialMap.toRationalMap_surjective f + trans (f.toRationalMap.toPartialMap.restrict _ + f.dense_domain f.le_domain_toRationalMap).toRationalMap + · simp + · congr 1 + exact PartialMap.ext _ f rfl (by simpa using f.toPartialMap_toRationalMap_restrict) + +instance [IsReduced X] [Y.IsSeparated] [S.IsSeparated] [X.Over S] [Y.Over S] + (f : X ⤏ Y) [f.IsOver S] : f.toPartialMap.IsOver S := by + rw [← PartialMap.isOver_toRationalMap_iff_of_isSeparated, f.toRationalMap_toPartialMap] + infer_instance + +end domain + end Scheme end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Restrict.lean b/Mathlib/AlgebraicGeometry/Restrict.lean index e1e6db017d998..4977b2dfaeffc 100644 --- a/Mathlib/AlgebraicGeometry/Restrict.lean +++ b/Mathlib/AlgebraicGeometry/Restrict.lean @@ -52,8 +52,10 @@ def ι : ↑U ⟶ X := X.ofRestrict _ instance : IsOpenImmersion U.ι := inferInstanceAs (IsOpenImmersion (X.ofRestrict _)) -@[simps] instance : U.toScheme.Over X := ⟨U.ι⟩ @[simps! over] instance : U.toScheme.CanonicallyOver X where + hom := U.ι + +instance (U : X.Opens) : U.ι.IsOver X where lemma toScheme_carrier : (U : Type u) = (U : Set X) := rfl @@ -202,6 +204,8 @@ lemma Scheme.homOfLE_ι (X : Scheme.{u}) {U V : X.Opens} (e : U ≤ V) : X.homOfLE e ≫ V.ι = U.ι := IsOpenImmersion.lift_fac _ _ _ +instance {U V : X.Opens} (h : U ≤ V) : (X.homOfLE h).IsOver X where + @[simp] lemma Scheme.homOfLE_rfl (X : Scheme.{u}) (U : X.Opens) : X.homOfLE (refl U) = 𝟙 _ := by rw [← cancel_mono U.ι, Scheme.homOfLE_ι, Category.id_comp] @@ -444,6 +448,21 @@ theorem isPullback_morphismRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Open -- Porting note: changed `rw` to `erw` erw [pullbackRestrictIsoRestrict_inv_fst]; rw [Category.comp_id] +lemma isPullback_opens_inf_le {X : Scheme} {U V W : X.Opens} (hU : U ≤ W) (hV : V ≤ W) : + IsPullback (X.homOfLE inf_le_left) (X.homOfLE inf_le_right) (X.homOfLE hU) (X.homOfLE hV) := by + refine (isPullback_morphismRestrict (X.homOfLE hV) (W.ι ⁻¹ᵁ U)).of_iso (V.ι.isoImage _ ≪≫ + X.isoOfEq ?_) (W.ι.isoImage _ ≪≫ X.isoOfEq ?_) (Iso.refl _) (Iso.refl _) ?_ ?_ ?_ ?_ + · rw [← TopologicalSpace.Opens.map_comp_obj, ← Scheme.comp_base, Scheme.homOfLE_ι] + exact V.functor_map_eq_inf U + · exact (W.functor_map_eq_inf U).trans (by simpa) + all_goals { simp [← cancel_mono (Scheme.Opens.ι _)] } + +lemma isPullback_opens_inf {X : Scheme} (U V : X.Opens) : + IsPullback (X.homOfLE inf_le_left) (X.homOfLE inf_le_right) U.ι V.ι := + (isPullback_morphismRestrict V.ι U).of_iso (V.ι.isoImage _ ≪≫ X.isoOfEq + (V.functor_map_eq_inf U)) (Iso.refl _) (Iso.refl _) (Iso.refl _) (by simp [← cancel_mono U.ι]) + (by simp [← cancel_mono V.ι]) (by simp) (by simp) + @[simp] lemma morphismRestrict_id {X : Scheme.{u}} (U : X.Opens) : 𝟙 X ∣_ U = 𝟙 _ := by rw [← cancel_mono U.ι, morphismRestrict_ι, Category.comp_id, Category.id_comp] diff --git a/Mathlib/CategoryTheory/Comma/OverClass.lean b/Mathlib/CategoryTheory/Comma/OverClass.lean index bfe80651c0003..c9ad20a1fd78e 100644 --- a/Mathlib/CategoryTheory/Comma/OverClass.lean +++ b/Mathlib/CategoryTheory/Comma/OverClass.lean @@ -31,7 +31,7 @@ universe v u variable {C : Type u} [Category.{v} C] -variable {X Y : C} (f : X ⟶ Y) (S S' : C) +variable {X Y Z : C} (f : X ⟶ Y) (S S' : C) /-- `OverClass X S` is the typeclass containing the data of a structure morphism `X ↘ S : X ⟶ S`. @@ -68,27 +68,39 @@ def CanonicallyOverClass.Simps.over (X S : C) [CanonicallyOverClass X S] : X ⟶ initialize_simps_projections CanonicallyOverClass (hom → over) @[simps] -instance (priority := 100) : OverClass X X := ⟨𝟙 _⟩ +instance : OverClass X X := ⟨𝟙 _⟩ -@[simps] +instance : IsIso (S ↘ S) := inferInstanceAs (IsIso (𝟙 S)) + +-- This cannot be a simp lemma be cause it loops with `comp_over`. +@[simps (config := .lemmasOnly)] instance (priority := 900) [CanonicallyOverClass X Y] [OverClass Y S] : OverClass X S := ⟨X ↘ Y ≫ Y ↘ S⟩ /-- Given `OverClass X S` and `OverClass Y S` and `f : X ⟶ Y`, `HomIsOver f S` is the typeclass asserting `f` commutes with the structure morphisms. -/ class HomIsOver (f : X ⟶ Y) (S : C) [OverClass X S] [OverClass Y S] : Prop where - comp_over : f ≫ Y ↘ S = X ↘ S := by simp + comp_over : f ≫ Y ↘ S = X ↘ S := by aesop @[reassoc (attr := simp)] lemma comp_over [OverClass X S] [OverClass Y S] [HomIsOver f S] : f ≫ Y ↘ S = X ↘ S := HomIsOver.comp_over +instance [OverClass X S] : HomIsOver (𝟙 X) S where + +instance [OverClass X S] [OverClass Y S] [OverClass Z S] + (f : X ⟶ Y) (g : Y ⟶ Z) [HomIsOver f S] [HomIsOver g S] : + HomIsOver (f ≫ g) S where + /-- `Scheme.IsOverTower X Y S` is the typeclass asserting that the structure morphisms `X ↘ Y`, `Y ↘ S`, and `X ↘ S` commute. -/ abbrev IsOverTower (X Y S : C) [OverClass X S] [OverClass Y S] [OverClass X Y] := HomIsOver (X ↘ Y) S +instance [OverClass X S] : IsOverTower X X S where +instance [OverClass X S] : IsOverTower X S S where + instance [CanonicallyOverClass X Y] [OverClass Y S] : IsOverTower X Y S := ⟨rfl⟩ From de26fff24458b9bcab999c26f961ecbc29f36c2d Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Sat, 23 Nov 2024 20:25:40 +0000 Subject: [PATCH 251/829] feat(NumberTheory/LSeries/PrimesInAP): more properties of von Mangoldt on residue classes (#19418) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This continues the sequence of PRs on the way to **Dirichlet's Theorem**. This one adds some properties of the von Mangoldt function `Λ` restricted to a residue class: if further restricted to primes, its support is the set of primes in that residue class, and the function `n ↦ Λ n / n`, restriced to non-primes in a residue class, is summable (the latter will be important in deducing information on the former). See [here](https://leanprover.zulipchat.com/#narrow/channel/144837-PR-reviews/topic/Prerequisites.20for.20PNT.20and.20Dirichlet's.20Thm/near/483904451) on Zulip. --- Mathlib/NumberTheory/LSeries/PrimesInAP.lean | 102 +++++++++++++++++++ 1 file changed, 102 insertions(+) diff --git a/Mathlib/NumberTheory/LSeries/PrimesInAP.lean b/Mathlib/NumberTheory/LSeries/PrimesInAP.lean index 0740c4659750a..6a882afa27f44 100644 --- a/Mathlib/NumberTheory/LSeries/PrimesInAP.lean +++ b/Mathlib/NumberTheory/LSeries/PrimesInAP.lean @@ -109,6 +109,108 @@ lemma abscissaOfAbsConv_residueClass_le_one : · simp +contextual only [term, Set.mem_setOf_eq, hn, not_false_eq_true, Set.indicator_of_not_mem, ofReal_zero, zero_div, ite_self] +/-- The set we are interested in (prime numbers in the residue class `a`) is the same as the support +of `ArithmeticFunction.vonMangoldt.residueClass` restricted to primes (and divided by `n`; +this is how this result is used later). -/ +lemma support_residueClass_prime_div : + Function.support (fun n : ℕ ↦ (if n.Prime then residueClass a n else 0) / n) = + {p : ℕ | p.Prime ∧ (p : ZMod q) = a} := by + simp only [Function.support, ne_eq, div_eq_zero_iff, ite_eq_right_iff, + Set.indicator_apply_eq_zero, Set.mem_setOf_eq, Nat.cast_eq_zero, not_or, Classical.not_imp] + ext1 p + simp only [Set.mem_setOf_eq] + exact ⟨fun H ↦ ⟨H.1.1, H.1.2.1⟩, + fun H ↦ ⟨⟨H.1, H.2, vonMangoldt_ne_zero_iff.mpr H.1.isPrimePow⟩, H.1.ne_zero⟩⟩ + +private noncomputable def F₀ (n : ℕ) : ℝ := (if n.Prime then 0 else vonMangoldt n) / n + +private noncomputable def F' (pk : Nat.Primes × ℕ) : ℝ := F₀ (pk.1 ^ (pk.2 + 1)) + +private noncomputable def F'' : Nat.Primes × ℕ → ℝ := F' ∘ (Prod.map _root_.id (· + 1)) + +private lemma F''_le (p : Nat.Primes) (k : ℕ) : F'' (p, k) ≤ 2 * (p : ℝ)⁻¹ ^ (k + 3 / 2 : ℝ) := + calc _ + _ = Real.log p * (p : ℝ)⁻¹ ^ (k + 2) := by + simp only [F'', Function.comp_apply, F', F₀, Prod.map_apply, id_eq, le_add_iff_nonneg_left, + zero_le, Nat.Prime.not_prime_pow, ↓reduceIte, vonMangoldt_apply_prime p.prop, + vonMangoldt_apply_pow (Nat.zero_ne_add_one _).symm, Nat.cast_pow, div_eq_mul_inv, + inv_pow (p : ℝ) (k + 2)] + _ ≤ (p: ℝ) ^ (1 / 2 : ℝ) / (1 / 2) * (p : ℝ)⁻¹ ^ (k + 2) := + mul_le_mul_of_nonneg_right (Real.log_le_rpow_div p.val.cast_nonneg one_half_pos) + (pow_nonneg (inv_nonneg_of_nonneg (Nat.cast_nonneg ↑p)) (k + 2)) + _ = 2 * (p : ℝ)⁻¹ ^ (-1 / 2 : ℝ) * (p : ℝ)⁻¹ ^ (k + 2) := by + simp only [← div_mul, div_one, mul_comm, neg_div, Real.inv_rpow p.val.cast_nonneg, + ← Real.rpow_neg p.val.cast_nonneg, neg_neg] + _ = _ := by + rw [mul_assoc, ← Real.rpow_natCast, + ← Real.rpow_add <| by have := p.prop.pos; positivity, Nat.cast_add, Nat.cast_two, + add_comm, add_assoc] + norm_num + +open Nat.Primes + +private lemma summable_F'' : Summable F'' := by + have hp₀ (p : Nat.Primes) : 0 < (p : ℝ)⁻¹ := inv_pos_of_pos (Nat.cast_pos.mpr p.prop.pos) + have hp₁ (p : Nat.Primes) : (p : ℝ)⁻¹ < 1 := + (inv_lt_one₀ <| mod_cast p.prop.pos).mpr <| Nat.one_lt_cast.mpr <| p.prop.one_lt + suffices Summable fun (pk : Nat.Primes × ℕ) ↦ (pk.1 : ℝ)⁻¹ ^ (pk.2 + 3 / 2 : ℝ) by + refine (Summable.mul_left 2 this).of_nonneg_of_le (fun pk ↦ ?_) (fun pk ↦ F''_le pk.1 pk.2) + simp only [F'', Function.comp_apply, F', F₀, Prod.map_fst, id_eq, Prod.map_snd, Nat.cast_pow] + have := vonMangoldt_nonneg (n := (pk.1 : ℕ) ^ (pk.2 + 2)) + positivity + conv => enter [1, pk]; rw [Real.rpow_add <| hp₀ pk.1, Real.rpow_natCast] + refine (summable_prod_of_nonneg (fun _ ↦ by positivity)).mpr ⟨(fun p ↦ ?_), ?_⟩ + · dsimp only -- otherwise the `exact` below times out + exact Summable.mul_right _ <| summable_geometric_of_lt_one (hp₀ p).le (hp₁ p) + · dsimp only + conv => enter [1, p]; rw [tsum_mul_right, tsum_geometric_of_lt_one (hp₀ p).le (hp₁ p)] + refine (summable_rpow.mpr (by norm_num : -(3 / 2 : ℝ) < -1)).mul_left 2 + |>.of_nonneg_of_le (fun p ↦ ?_) (fun p ↦ ?_) + · have := sub_pos.mpr (hp₁ p) + positivity + · rw [Real.inv_rpow p.val.cast_nonneg, Real.rpow_neg p.val.cast_nonneg] + gcongr + rw [inv_le_comm₀ (sub_pos.mpr (hp₁ p)) zero_lt_two, le_sub_comm, + show (1 : ℝ) - 2⁻¹ = 2⁻¹ by norm_num, inv_le_inv₀ (mod_cast p.prop.pos) zero_lt_two] + exact Nat.ofNat_le_cast.mpr p.prop.two_le + +/-- The function `n ↦ Λ n / n`, restriced to non-primes in a residue class, is summable. +This is used to convert results on `ArithmeticFunction.vonMangoldt.residueClass` to results +on primes in an arithmetic progression. -/ +lemma summable_residueClass_non_primes_div : + Summable fun n : ℕ ↦ (if n.Prime then 0 else residueClass a n) / n := by + have h₀ (n : ℕ) : 0 ≤ (if n.Prime then 0 else residueClass a n) / n := by + have := residueClass_nonneg a n + positivity + have hleF₀ (n : ℕ) : (if n.Prime then 0 else residueClass a n) / n ≤ F₀ n := by + refine div_le_div_of_nonneg_right ?_ n.cast_nonneg + split_ifs; exacts [le_rfl, residueClass_le a n] + refine Summable.of_nonneg_of_le h₀ hleF₀ ?_ + have hF₀ (p : Nat.Primes) : F₀ p.val = 0 := by + simp only [p.prop, ↓reduceIte, zero_div, F₀] + refine (summable_subtype_iff_indicator (s := {n | IsPrimePow n}).mp ?_).congr + fun n ↦ Set.indicator_apply_eq_self.mpr fun (hn : ¬ IsPrimePow n) ↦ ?_ + swap + · simp +contextual only [div_eq_zero_iff, ite_eq_left_iff, vonMangoldt_eq_zero_iff, hn, + not_false_eq_true, implies_true, Nat.cast_eq_zero, true_or, F₀] + have hFF' : + F₀ ∘ Subtype.val (p := fun n ↦ n ∈ {n | IsPrimePow n}) = F' ∘ ⇑prodNatEquiv.symm := by + refine (Equiv.eq_comp_symm prodNatEquiv (F₀ ∘ Subtype.val) F').mpr ?_ + ext1 n + simp only [Function.comp_apply, F'] + congr + rw [hFF'] + refine (Nat.Primes.prodNatEquiv.symm.summable_iff (f := F')).mpr ?_ + have hF'₀ (p : Nat.Primes) : F' (p, 0) = 0 := by simp only [zero_add, pow_one, hF₀, F'] + have hF'₁ : F'' = F' ∘ (Prod.map _root_.id (· + 1)) := by + ext1 + simp only [Function.comp_apply, Prod.map_fst, id_eq, Prod.map_snd, F'', F'] + refine (Function.Injective.summable_iff ?_ fun u hu ↦ ?_).mp <| hF'₁ ▸ summable_F'' + · exact Function.Injective.prodMap (fun ⦃a₁ a₂⦄ a ↦ a) <| add_left_injective 1 + · simp only [Set.range_prod_map, Set.range_id, Set.mem_prod, Set.mem_univ, Set.mem_range, + Nat.exists_add_one_eq, true_and, not_lt, nonpos_iff_eq_zero] at hu + rw [← hF'₀ u.1, ← hu] + variable [NeZero q] {a} /-- We can express `ArithmeticFunction.vonMangoldt.residueClass` as a linear combination From 9b9e6ffdfca324d29cf397ddf8a9cb8151a04240 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Sat, 23 Nov 2024 20:34:16 +0000 Subject: [PATCH 252/829] feat(NumberTheory/Padics): Mahler coeffs tend to 0 (#19340) This adds the key technical lemma in the proof of Mahler's theorem characterising continuous functions on Zp: for any continuous function, the iterated forward differences at 0 tend to zero. Co-authored-by: Giulio Caflisch --- Mathlib/NumberTheory/Padics/MahlerBasis.lean | 149 ++++++++++++++++++- docs/references.bib | 13 ++ 2 files changed, 159 insertions(+), 3 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/MahlerBasis.lean b/Mathlib/NumberTheory/Padics/MahlerBasis.lean index 543eaf0533d7d..a6be858db5335 100644 --- a/Mathlib/NumberTheory/Padics/MahlerBasis.lean +++ b/Mathlib/NumberTheory/Padics/MahlerBasis.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 David Loeffler. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Giulio Caflisch, David Loeffler -/ +import Mathlib.Algebra.Group.ForwardDiff import Mathlib.Analysis.Normed.Group.Ultra import Mathlib.NumberTheory.Padics.ProperSpace import Mathlib.RingTheory.Binomial @@ -15,15 +16,26 @@ import Mathlib.Topology.ContinuousMap.Compact In this file we introduce the Mahler basis function `mahler k`, for `k : ℕ`, which is the unique continuous map `ℤ_[p] → ℚ_[p]` agreeing with `n ↦ n.choose k` for `n ∈ ℕ`. -In future PR's, we will show that these functions give a Banach basis for the space of continuous -maps `ℤ_[p] → ℚ_[p]`. +We also show that for any continuous function `f` on `ℤ_[p]` (valued in a `p`-adic normed space), +the iterated forward differences `Δ^[n] f 0` tend to 0. For this, we follow the argument of +Bojanić [bojanic74]. + +In future PR's, we will show that the Mahler functions give a Banach basis for the space of +continuous maps `ℤ_[p] → ℚ_[p]`, with the basis coefficients of `f` given by the forward differences +`Δ^[n] f 0`. ## References +* [R. Bojanić, *A simple proof of Mahler's theorem on approximation of continuous functions of a + p-adic variable by polynomials*][bojanic74] * [P. Colmez, *Fonctions d'une variable p-adique*][colmez2010] + +## Tags + +Bojanic -/ -open Finset +open Finset fwdDiff IsUltrametricDist NNReal Filter Topology variable {p : ℕ} [hp : Fact p.Prime] @@ -95,3 +107,134 @@ The uniform norm of the `k`-th Mahler basis function is 1, for every `k`. · -- Show norm 1 is attained at `x = k` refine (le_of_eq ?_).trans ((mahler k).norm_coe_le_norm k) rw [mahler_natCast_eq, Nat.choose_self, Nat.cast_one, norm_one] + +section fwdDiff + +variable {M G : Type*} + +/-- Bound for iterated forward differences of a continuous function from a compact space to a +nonarchimedean seminormed group. -/ +lemma IsUltrametricDist.norm_fwdDiff_iter_apply_le [TopologicalSpace M] [CompactSpace M] + [AddCommMonoid M] [SeminormedAddCommGroup G] [IsUltrametricDist G] + (h : M) (f : C(M, G)) (m : M) (n : ℕ) : ‖Δ_[h]^[n] f m‖ ≤ ‖f‖ := by + -- A proof by induction on `n` would be possible but would involve some messing around to + -- define `Δ_[h]` as an operator on continuous maps (not just on bare functions). So instead we + -- use the formula for `Δ_[h]^[n] f` as a sum. + rw [fwdDiff_iter_eq_sum_shift] + refine norm_sum_le_of_forall_le_of_nonneg (norm_nonneg f) fun i _ ↦ ?_ + exact (norm_zsmul_le _ _).trans (f.norm_coe_le_norm _) + +/-- First step in Bojanić's proof of Mahler's theorem (equation (10) of [bojanic74]): rewrite +`Δ^[n + R] f 0` in a shape that makes it easy to bound `p`-adically. -/ +private lemma bojanic_mahler_step1 [AddCommMonoidWithOne M] [AddCommGroup G] (f : M → G) + (n : ℕ) {R : ℕ} (hR : 1 ≤ R) : + Δ_[1]^[n + R] f 0 = -∑ j ∈ range (R - 1), R.choose (j + 1) • Δ_[1]^[n + (j + 1)] f 0 + + ∑ k ∈ range (n + 1), ((-1 : ℤ) ^ (n - k) * n.choose k) • (f (k + R) - f k) := by + have aux : Δ_[1]^[n + R] f 0 = R.choose (R - 1 + 1) • Δ_[1]^[n + R] f 0 := by + rw [Nat.sub_add_cancel hR, Nat.choose_self, one_smul] + rw [neg_add_eq_sub, eq_sub_iff_add_eq, add_comm, aux, (by omega : n + R = (n + ((R - 1) + 1))), + ← sum_range_succ, Nat.sub_add_cancel hR, + ← sub_eq_iff_eq_add.mpr (sum_range_succ' (fun x ↦ R.choose x • Δ_[1]^[n + x] f 0) R), add_zero, + Nat.choose_zero_right, one_smul] + have : ∑ k ∈ Finset.range (R + 1), R.choose k • Δ_[1]^[n + k] f 0 = Δ_[1]^[n] f R := by + simpa only [← Function.iterate_add_apply, add_comm, nsmul_one, add_zero] using + (shift_eq_sum_fwdDiff_iter 1 (Δ_[1]^[n] f) R 0).symm + simp only [this, fwdDiff_iter_eq_sum_shift (1 : M) f n, mul_comm, nsmul_one, mul_smul, add_comm, + add_zero, smul_sub, sum_sub_distrib] + +end fwdDiff + +namespace PadicInt + +section norm_fwdDiff + +variable {p : ℕ} [hp : Fact p.Prime] {E : Type*} + [NormedAddCommGroup E] [NormedSpace ℚ_[p] E] [IsUltrametricDist E] + +/-- +Second step in Bojanić's proof of Mahler's theorem (equation (11) of [bojanic74]): show that values +`Δ_[1]^[n + p ^ t] f 0` for large enough `n` are bounded by the max of `(‖f‖ / p ^ s)` and `1 / p` +times a sup over values for smaller `n`. + +We use `nnnorm`s on the RHS since `Finset.sup` requires an order with a bottom element. +-/ +private lemma bojanic_mahler_step2 {f : C(ℤ_[p], E)} {s t : ℕ} + (hst : ∀ x y : ℤ_[p], ‖x - y‖ ≤ p ^ (-t : ℤ) → ‖f x - f y‖ ≤ ‖f‖ / p ^ s) (n : ℕ) : + ‖Δ_[1]^[n + p ^ t] f 0‖ ≤ max ↑((Finset.range (p ^ t - 1)).sup + fun j ↦ ‖Δ_[1]^[n + (j + 1)] f 0‖₊ / p) (‖f‖ / p ^ s) := by + -- Use previous lemma to rewrite in a convenient form. + rw [bojanic_mahler_step1 _ _ (one_le_pow₀ hp.out.one_le)] + -- Now use ultrametric property and bound each term separately. + refine (norm_add_le_max _ _).trans (max_le_max ?_ ?_) + · -- Bounding the sum over `range (p ^ t - 1)`: every term involves a value `Δ_[1]^[·] f 0` and + -- a binomial coefficient which is divisible by `p` + rw [norm_neg, ← coe_nnnorm, coe_le_coe] + refine nnnorm_sum_le_of_forall_le (fun i hi ↦ Finset.le_sup_of_le hi ?_) + rw [mem_range] at hi + rw [← Nat.cast_smul_eq_nsmul ℚ_[p], nnnorm_smul, div_eq_inv_mul] + refine mul_le_mul_of_nonneg_right ?_ (by simp only [zero_le]) + -- remains to show norm of binomial coeff is `≤ p⁻¹` + have : 0 < (p ^ t).choose (i + 1) := Nat.choose_pos (by omega) + rw [← zpow_neg_one, ← coe_le_coe, coe_nnnorm, Padic.norm_eq_pow_val (mod_cast this.ne'), + coe_zpow, NNReal.coe_natCast, (zpow_right_strictMono₀ (mod_cast hp.out.one_lt)).le_iff_le, + neg_le_neg_iff, Padic.valuation_natCast, Nat.one_le_cast] + exact one_le_padicValNat_of_dvd this <| hp.out.dvd_choose_pow (by omega) (by omega) + · -- Bounding the sum over `range (n + 1)`: every term is small by the choice of `t` + refine norm_sum_le_of_forall_le_of_nonempty nonempty_range_succ (fun i _ ↦ ?_) + calc ‖((-1 : ℤ) ^ (n - i) * n.choose i) • (f (i + ↑(p ^ t)) - f i)‖ + _ ≤ ‖f (i + ↑(p ^ t)) - f i‖ := by + rw [← Int.cast_smul_eq_zsmul ℚ_[p], norm_smul] + apply mul_le_of_le_one_left (norm_nonneg _) + simpa only [← coe_intCast] using norm_le_one _ + _ ≤ ‖f‖ / p ^ s := by + apply hst + rw [Nat.cast_pow, add_sub_cancel_left, norm_pow, norm_p, inv_pow, zpow_neg, zpow_natCast] + +/-- +Explicit bound for the decay rate of the Mahler coefficients of a continuous function on `ℤ_[p]`. +This will be used to prove Mahler's theorem. + -/ +lemma fwdDiff_iter_le_of_forall_le {f : C(ℤ_[p], E)} {s t : ℕ} + (hst : ∀ x y : ℤ_[p], ‖x - y‖ ≤ p ^ (-t : ℤ) → ‖f x - f y‖ ≤ ‖f‖ / p ^ s) (n : ℕ) : + ‖Δ_[1]^[n + s * p ^ t] f 0‖ ≤ ‖f‖ / p ^ s := by + -- We show the following more general statement by induction on `k`: + suffices ∀ {k : ℕ}, k ≤ s → ‖Δ_[1]^[n + k * p ^ t] f 0‖ ≤ ‖f‖ / p ^ k from this le_rfl + intro k hk + induction' k with k IH generalizing n + · -- base case just says that `‖Δ^[·] (⇑f) 0‖` is bounded by `‖f‖` + simpa only [zero_mul, pow_zero, add_zero, div_one] using norm_fwdDiff_iter_apply_le 1 f 0 n + · -- induction is the "step 2" lemma above + rw [add_mul, one_mul, ← add_assoc] + refine (bojanic_mahler_step2 hst (n + k * p ^ t)).trans (max_le ?_ ?_) + · rw [← coe_nnnorm, ← NNReal.coe_natCast, ← NNReal.coe_pow, ← NNReal.coe_div, NNReal.coe_le_coe] + refine Finset.sup_le fun j _ ↦ ?_ + rw [pow_succ, ← div_div, div_le_div_iff_of_pos_right (mod_cast hp.out.pos), add_right_comm] + exact_mod_cast IH (n + (j + 1)) (by omega) + · exact div_le_div_of_nonneg_left (norm_nonneg _) + (mod_cast pow_pos hp.out.pos _) (mod_cast pow_le_pow_right₀ hp.out.one_le hk) + +/-- Key lemma for Mahler's theorem: for `f` a continuous function on `ℤ_[p]`, the sequence +`n ↦ Δ^[n] f 0` tends to 0. See `PadicInt.fwdDiff_iter_le_of_forall_le` for an explicit +estimate of the decay rate. -/ +lemma fwdDiff_tendsto_zero (f : C(ℤ_[p], E)) : Tendsto (Δ_[1]^[·] f 0) atTop (𝓝 0) := by + -- first extract an `s` + refine NormedAddCommGroup.tendsto_nhds_zero.mpr (fun ε hε ↦ ?_) + have : Tendsto (fun s ↦ ‖f‖ / p ^ s) _ _ := tendsto_const_nhds.div_atTop + (tendsto_pow_atTop_atTop_of_one_lt (mod_cast hp.out.one_lt)) + obtain ⟨s, hs⟩ := (this.eventually_lt_const hε).exists + refine .mp ?_ (.of_forall fun x hx ↦ lt_of_le_of_lt hx hs) + -- use uniform continuity to find `t` + obtain ⟨t, ht⟩ : ∃ t : ℕ, ∀ x y, ‖x - y‖ ≤ p ^ (-t : ℤ) → ‖f x - f y‖ ≤ ‖f‖ / p ^ s := by + rcases eq_or_ne f 0 with rfl | hf + · -- silly case : f = 0 + simp + have : 0 < ‖f‖ / p ^ s := div_pos (norm_pos_iff.mpr hf) (mod_cast pow_pos hp.out.pos _) + obtain ⟨δ, hδpos, hδf⟩ := f.uniform_continuity _ this + obtain ⟨t, ht⟩ := PadicInt.exists_pow_neg_lt p hδpos + exact ⟨t, fun x y hxy ↦ by simpa only [dist_eq_norm_sub] using (hδf (hxy.trans_lt ht)).le⟩ + filter_upwards [eventually_ge_atTop (s * p ^ t)] with m hm + simpa only [Nat.sub_add_cancel hm] using fwdDiff_iter_le_of_forall_le ht (m - s * p ^ t) + +end norm_fwdDiff + +end PadicInt diff --git a/docs/references.bib b/docs/references.bib index f227b4270fd17..3d890cec41f2a 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -423,6 +423,19 @@ @Book{ bogachev2007 url = {https://doi.org/10.1007/978-3-540-34514-5} } +@Article{ bojanic74, + author = {Bojani\'c, Ranko}, + title = {A simple proof of {M}ahler's theorem on approximation of + continuous functions of a p-adic variable by polynomials}, + journal = {Journal of Number Theory}, + volume = {6}, + pages = {412-415}, + year = {1974}, + issn = {0022-314X}, + doi = {10.1016/0022-314X(74)90038-9}, + url = {https://doi.org/10.1016/0022-314X(74)90038-9} +} + @Book{ bollobas1986, author = {Bollob\'{a}s, B\'{e}la}, title = {Combinatorics: Set Systems, Hypergraphs, Families of From b5155f3e85865ffe02cfd495803e84e98ddab21f Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 23 Nov 2024 22:52:24 +0000 Subject: [PATCH 253/829] chore: more simpNF cleanup (#19395) --- Mathlib/LinearAlgebra/Multilinear/Basic.lean | 24 +------------------ Mathlib/NumberTheory/Modular.lean | 3 --- .../ModularForms/SlashActions.lean | 6 ++++- .../NumberTheory/Padics/PadicIntegers.lean | 16 ++++++------- Mathlib/NumberTheory/Padics/RingHoms.lean | 2 +- Mathlib/Order/Closure.lean | 2 +- Mathlib/Order/Compare.lean | 6 ++--- Mathlib/Order/Filter/Germ/Basic.lean | 2 +- Mathlib/Order/Filter/Pointwise.lean | 19 ++++----------- Mathlib/Order/Heyting/Basic.lean | 11 ++++----- Mathlib/Order/Interval/Basic.lean | 1 - Mathlib/Order/LiminfLimsup.lean | 6 ++--- Mathlib/Order/SupIndep.lean | 14 +---------- .../RingTheory/DedekindDomain/SInteger.lean | 1 - Mathlib/RingTheory/Ideal/Operations.lean | 4 +--- Mathlib/RingTheory/Valuation/Basic.lean | 2 +- .../Valuation/ValuationSubring.lean | 8 ++----- 17 files changed, 35 insertions(+), 92 deletions(-) diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index bef753458eb59..96de9639e3cfd 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -1719,7 +1719,6 @@ theorem curryFinFinset_symm_apply {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = m <| finSumEquivOfFinset hk hl (Sum.inr i) := rfl --- @[simp] -- Porting note: simpNF linter, lhs simplifies, added aux version below theorem curryFinFinset_symm_apply_piecewise_const {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = k) (hl : #sᶜ = l) (f : MultilinearMap R (fun _ : Fin k => M') (MultilinearMap R (fun _ : Fin l => M') M₂)) @@ -1734,18 +1733,6 @@ theorem curryFinFinset_symm_apply_piecewise_const {k l n : ℕ} {s : Finset (Fin rw [finSumEquivOfFinset_inr, Finset.piecewise_eq_of_not_mem] exact Finset.mem_compl.1 (Finset.orderEmbOfFin_mem _ _ _) -@[simp] -theorem curryFinFinset_symm_apply_piecewise_const_aux {k l n : ℕ} {s : Finset (Fin n)} - (hk : #s = k) (hl : #sᶜ = l) - (f : MultilinearMap R (fun _ : Fin k => M') (MultilinearMap R (fun _ : Fin l => M') M₂)) - (x y : M') : - ((⇑f fun _ => x) (fun i => (Finset.piecewise s (fun _ => x) (fun _ => y) - ((sᶜ.orderEmbOfFin hl) i))) = f (fun _ => x) fun _ => y) := by - have := curryFinFinset_symm_apply_piecewise_const hk hl f x y - simp only [curryFinFinset_symm_apply, finSumEquivOfFinset_inl, Finset.orderEmbOfFin_mem, - Finset.piecewise_eq_of_mem, finSumEquivOfFinset_inr] at this - exact this - @[simp] theorem curryFinFinset_symm_apply_const {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = k) (hl : #sᶜ = l) @@ -1753,23 +1740,14 @@ theorem curryFinFinset_symm_apply_const {k l n : ℕ} {s : Finset (Fin n)} (hk : (x : M') : ((curryFinFinset R M₂ M' hk hl).symm f fun _ => x) = f (fun _ => x) fun _ => x := rfl --- @[simp] -- Porting note: simpNF, lhs simplifies, added aux version below theorem curryFinFinset_apply_const {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = k) (hl : #sᶜ = l) (f : MultilinearMap R (fun _ : Fin n => M') M₂) (x y : M') : (curryFinFinset R M₂ M' hk hl f (fun _ => x) fun _ => y) = f (s.piecewise (fun _ => x) fun _ => y) := by + -- Porting note: `rw` fails refine (curryFinFinset_symm_apply_piecewise_const hk hl _ _ _).symm.trans ?_ - -- `rw` fails rw [LinearEquiv.symm_apply_apply] -@[simp] -theorem curryFinFinset_apply_const_aux {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = k) - (hl : #sᶜ = l) (f : MultilinearMap R (fun _ : Fin n => M') M₂) (x y : M') : - (f fun i => Sum.elim (fun _ => x) (fun _ => y) ((⇑ (Equiv.symm (finSumEquivOfFinset hk hl))) i)) - = f (s.piecewise (fun _ => x) fun _ => y) := by - rw [← curryFinFinset_apply] - apply curryFinFinset_apply_const - end MultilinearMap end Currying diff --git a/Mathlib/NumberTheory/Modular.lean b/Mathlib/NumberTheory/Modular.lean index ff3eef6b03cda..7eb5dbb5ffec4 100644 --- a/Mathlib/NumberTheory/Modular.lean +++ b/Mathlib/NumberTheory/Modular.lean @@ -180,9 +180,6 @@ def lcRow0Extend {cd : Fin 2 → ℤ} (hcd : IsCoprime (cd 0) (cd 1)) : rw [neg_sq] exact hcd.sq_add_sq_ne_zero, LinearEquiv.refl ℝ (Fin 2 → ℝ)] --- `simpNF` times out, but only in CI where all of `Mathlib` is imported -attribute [nolint simpNF] lcRow0Extend_apply lcRow0Extend_symm_apply - /-- The map `lcRow0` is proper, that is, preimages of cocompact sets are finite in `[[* , *], [c, d]]`. -/ theorem tendsto_lcRow0 {cd : Fin 2 → ℤ} (hcd : IsCoprime (cd 0) (cd 1)) : diff --git a/Mathlib/NumberTheory/ModularForms/SlashActions.lean b/Mathlib/NumberTheory/ModularForms/SlashActions.lean index bc52891b5dffc..ef9d1a0bcb146 100644 --- a/Mathlib/NumberTheory/ModularForms/SlashActions.lean +++ b/Mathlib/NumberTheory/ModularForms/SlashActions.lean @@ -143,10 +143,14 @@ theorem is_invariant_const (A : SL(2, ℤ)) (x : ℂ) : zpow_neg, zpow_one, inv_one, mul_one, neg_zero, zpow_zero] /-- The constant function 1 is invariant under any element of `SL(2, ℤ)`. -/ --- @[simp] -- Porting note: simpNF says LHS simplifies to something more complex theorem is_invariant_one (A : SL(2, ℤ)) : (1 : ℍ → ℂ) ∣[(0 : ℤ)] A = (1 : ℍ → ℂ) := is_invariant_const _ _ +/-- Variant of `is_invariant_one` with the left hand side in simp normal form. -/ +@[simp] +theorem is_invariant_one' (A : SL(2, ℤ)) : (1 : ℍ → ℂ) ∣[(0 : ℤ)] (A : GL(2, ℝ)⁺) = 1 := by + simpa using is_invariant_one A + /-- A function `f : ℍ → ℂ` is slash-invariant, of weight `k ∈ ℤ` and level `Γ`, if for every matrix `γ ∈ Γ` we have `f(γ • z)= (c*z+d)^k f(z)` where `γ= ![![a, b], ![c, d]]`, and it acts on `ℍ` via Möbius transformations. -/ diff --git a/Mathlib/NumberTheory/Padics/PadicIntegers.lean b/Mathlib/NumberTheory/Padics/PadicIntegers.lean index 53e32586436a1..f9e0e0463f74b 100644 --- a/Mathlib/NumberTheory/Padics/PadicIntegers.lean +++ b/Mathlib/NumberTheory/Padics/PadicIntegers.lean @@ -153,8 +153,7 @@ def Coe.ringHom : ℤ_[p] →+* ℚ_[p] := (subring p).subtype @[simp, norm_cast] theorem coe_pow (x : ℤ_[p]) (n : ℕ) : (↑(x ^ n) : ℚ_[p]) = (↑x : ℚ_[p]) ^ n := rfl --- @[simp] -- Porting note: not in simpNF -theorem mk_coe (k : ℤ_[p]) : (⟨k, k.2⟩ : ℤ_[p]) = k := Subtype.coe_eta _ _ +theorem mk_coe (k : ℤ_[p]) : (⟨k, k.2⟩ : ℤ_[p]) = k := by simp /-- The inverse of a `p`-adic integer with norm equal to `1` is also a `p`-adic integer. Otherwise, the inverse is defined to be `0`. -/ @@ -165,10 +164,8 @@ instance : CharZero ℤ_[p] where cast_injective m n h := Nat.cast_injective (R := ℚ_[p]) (by rw [Subtype.ext_iff] at h; norm_cast at h) -@[norm_cast] -- @[simp] -- Porting note: not in simpNF -theorem intCast_eq (z1 z2 : ℤ) : (z1 : ℤ_[p]) = z2 ↔ z1 = z2 := by - suffices (z1 : ℚ_[p]) = z2 ↔ z1 = z2 from Iff.trans (by norm_cast) this - norm_cast +@[norm_cast] +theorem intCast_eq (z1 z2 : ℤ) : (z1 : ℤ_[p]) = z2 ↔ z1 = z2 := by simp @[deprecated (since := "2024-04-05")] alias coe_int_eq := intCast_eq @@ -280,8 +277,7 @@ theorem norm_eq_padic_norm {q : ℚ_[p]} (hq : ‖q‖ ≤ 1) : @norm ℤ_[p] _ @[simp] theorem norm_p : ‖(p : ℤ_[p])‖ = (p : ℝ)⁻¹ := padicNormE.norm_p --- @[simp] -- Porting note: not in simpNF -theorem norm_p_pow (n : ℕ) : ‖(p : ℤ_[p]) ^ n‖ = (p : ℝ) ^ (-n : ℤ) := padicNormE.norm_p_pow n +theorem norm_p_pow (n : ℕ) : ‖(p : ℤ_[p]) ^ n‖ = (p : ℝ) ^ (-n : ℤ) := by simp private def cauSeq_to_rat_cauSeq (f : CauSeq ℤ_[p] norm) : CauSeq ℚ_[p] fun a => ‖a‖ := ⟨fun n => f n, fun _ hε => by simpa [norm, norm_def] using f.cauchy hε⟩ @@ -413,10 +409,12 @@ theorem norm_lt_one_mul {z1 z2 : ℤ_[p]} (hz2 : ‖z2‖ < 1) : ‖z1 * z2‖ < ‖z1 * z2‖ = ‖z1‖ * ‖z2‖ := by simp _ < 1 := mul_lt_one_of_nonneg_of_lt_one_right (norm_le_one _) (norm_nonneg _) hz2 --- @[simp] -- Porting note: not in simpNF theorem mem_nonunits {z : ℤ_[p]} : z ∈ nonunits ℤ_[p] ↔ ‖z‖ < 1 := by rw [lt_iff_le_and_ne]; simp [norm_le_one z, nonunits, isUnit_iff] +theorem not_isUnit_iff {z : ℤ_[p]} : ¬IsUnit z ↔ ‖z‖ < 1 := by + simpa using mem_nonunits + /-- A `p`-adic number `u` with `‖u‖ = 1` is a unit of `ℤ_[p]`. -/ def mkUnits {u : ℚ_[p]} (h : ‖u‖ = 1) : ℤ_[p]ˣ := let z : ℤ_[p] := ⟨u, le_of_eq h⟩ diff --git a/Mathlib/NumberTheory/Padics/RingHoms.lean b/Mathlib/NumberTheory/Padics/RingHoms.lean index dab80d6e19bd6..b80610a74c52a 100644 --- a/Mathlib/NumberTheory/Padics/RingHoms.lean +++ b/Mathlib/NumberTheory/Padics/RingHoms.lean @@ -393,7 +393,7 @@ theorem ker_toZModPow (n : ℕ) : rw [zmod_congr_of_sub_mem_span n x _ 0 _ h, cast_zero] apply appr_spec --- @[simp] -- Porting note: not in simpNF +-- This is not a simp lemma; simp can't match the LHS. theorem zmod_cast_comp_toZModPow (m n : ℕ) (h : m ≤ n) : (ZMod.castHom (pow_dvd_pow p h) (ZMod (p ^ m))).comp (@toZModPow p _ n) = @toZModPow p _ m := by apply ZMod.ringHom_eq_of_ker_eq diff --git a/Mathlib/Order/Closure.lean b/Mathlib/Order/Closure.lean index b9c15b5db30cc..8fb5b69f0275c 100644 --- a/Mathlib/Order/Closure.lean +++ b/Mathlib/Order/Closure.lean @@ -387,7 +387,7 @@ variable [PartialOrder α] [PartialOrder β] {u : β → α} (l : LowerAdjoint u theorem mem_closed_iff_closure_le (x : α) : x ∈ l.closed ↔ u (l x) ≤ x := l.closureOperator.isClosed_iff_closure_le -@[simp, nolint simpNF] -- Porting note: lemma does prove itself, seems to be a linter error +@[simp] theorem closure_is_closed (x : α) : u (l x) ∈ l.closed := l.idempotent x diff --git a/Mathlib/Order/Compare.lean b/Mathlib/Order/Compare.lean index 9fd8b1100315f..5e29b6e29a069 100644 --- a/Mathlib/Order/Compare.lean +++ b/Mathlib/Order/Compare.lean @@ -137,8 +137,7 @@ theorem cmp_swap [Preorder α] [@DecidableRel α (· < ·)] (a b : α) : (cmp a by_cases h : a < b <;> by_cases h₂ : b < a <;> simp [h, h₂, Ordering.swap] exact lt_asymm h h₂ --- Porting note: Not sure why the simpNF linter doesn't like this. @semorrison -@[simp, nolint simpNF] +@[simp] theorem cmpLE_toDual [LE α] [@DecidableRel α (· ≤ ·)] (x y : α) : cmpLE (toDual x) (toDual y) = cmpLE y x := rfl @@ -148,8 +147,7 @@ theorem cmpLE_ofDual [LE α] [@DecidableRel α (· ≤ ·)] (x y : αᵒᵈ) : cmpLE (ofDual x) (ofDual y) = cmpLE y x := rfl --- Porting note: Not sure why the simpNF linter doesn't like this. @semorrison -@[simp, nolint simpNF] +@[simp] theorem cmp_toDual [LT α] [@DecidableRel α (· < ·)] (x y : α) : cmp (toDual x) (toDual y) = cmp y x := rfl diff --git a/Mathlib/Order/Filter/Germ/Basic.lean b/Mathlib/Order/Filter/Germ/Basic.lean index 21eda99605828..00156fd9037ce 100644 --- a/Mathlib/Order/Filter/Germ/Basic.lean +++ b/Mathlib/Order/Filter/Germ/Basic.lean @@ -237,7 +237,7 @@ theorem coe_compTendsto (f : α → β) {lc : Filter γ} {g : γ → α} (hg : T rfl -- Porting note https://github.com/leanprover-community/mathlib4/issues/10959 --- simp cannot prove this +-- simp can't match the LHS. @[simp, nolint simpNF] theorem compTendsto'_coe (f : Germ l β) {lc : Filter γ} {g : γ → α} (hg : Tendsto g lc l) : f.compTendsto' _ hg.germ_tendsto = f.compTendsto g hg := diff --git a/Mathlib/Order/Filter/Pointwise.lean b/Mathlib/Order/Filter/Pointwise.lean index c76d14970e05a..05f1a4c1cf4e0 100644 --- a/Mathlib/Order/Filter/Pointwise.lean +++ b/Mathlib/Order/Filter/Pointwise.lean @@ -146,9 +146,8 @@ theorem pureOneHom_apply (a : α) : pureOneHom a = pure a := variable [One β] @[to_additive] --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11119): removed `simp` attribute because `simpNF` says it can prove it. protected theorem map_one [FunLike F α β] [OneHomClass F α β] (φ : F) : map φ 1 = 1 := by - rw [Filter.map_one', map_one, pure_one] + simp end One @@ -303,9 +302,7 @@ theorem mul_pure : f * pure b = f.map (· * b) := map₂_pure_right @[to_additive] --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11119): removed `simp` attribute because `simpNF` says it can prove it. -theorem pure_mul_pure : (pure a : Filter α) * pure b = pure (a * b) := - map₂_pure +theorem pure_mul_pure : (pure a : Filter α) * pure b = pure (a * b) := by simp @[to_additive (attr := simp)] theorem le_mul_iff : h ≤ f * g ↔ ∀ ⦃s⦄, s ∈ f → ∀ ⦃t⦄, t ∈ g → s * t ∈ h := @@ -408,9 +405,7 @@ theorem div_pure : f / pure b = f.map (· / b) := map₂_pure_right @[to_additive] --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11119): removed `simp` attribute because `simpNF` says it can prove it. -theorem pure_div_pure : (pure a : Filter α) / pure b = pure (a / b) := - map₂_pure +theorem pure_div_pure : (pure a : Filter α) / pure b = pure (a / b) := by simp @[to_additive] protected theorem div_le_div : f₁ ≤ f₂ → g₁ ≤ g₂ → f₁ / g₁ ≤ f₂ / g₂ := @@ -826,9 +821,7 @@ theorem smul_pure : f • pure b = f.map (· • b) := map₂_pure_right @[to_additive] --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11119): removed `simp` attribute because `simpNF` says it can prove it. -theorem pure_smul_pure : (pure a : Filter α) • (pure b : Filter β) = pure (a • b) := - map₂_pure +theorem pure_smul_pure : (pure a : Filter α) • (pure b : Filter β) = pure (a • b) := by simp @[to_additive] theorem smul_le_smul : f₁ ≤ f₂ → g₁ ≤ g₂ → f₁ • g₁ ≤ f₂ • g₂ := @@ -914,9 +907,7 @@ theorem pure_vsub : (pure a : Filter β) -ᵥ g = g.map (a -ᵥ ·) := theorem vsub_pure : f -ᵥ pure b = f.map (· -ᵥ b) := map₂_pure_right --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11119): removed `simp` attribute because `simpNF` says it can prove it. -theorem pure_vsub_pure : (pure a : Filter β) -ᵥ pure b = (pure (a -ᵥ b) : Filter α) := - map₂_pure +theorem pure_vsub_pure : (pure a : Filter β) -ᵥ pure b = (pure (a -ᵥ b) : Filter α) := by simp theorem vsub_le_vsub : f₁ ≤ f₂ → g₁ ≤ g₂ → f₁ -ᵥ g₁ ≤ f₂ -ᵥ g₂ := map₂_mono diff --git a/Mathlib/Order/Heyting/Basic.lean b/Mathlib/Order/Heyting/Basic.lean index 4c1f777c64b8f..c44de8d3c3d14 100644 --- a/Mathlib/Order/Heyting/Basic.lean +++ b/Mathlib/Order/Heyting/Basic.lean @@ -1084,11 +1084,11 @@ theorem top_eq : (⊤ : PUnit) = unit := theorem bot_eq : (⊥ : PUnit) = unit := rfl -@[simp, nolint simpNF] +@[simp] theorem sup_eq : a ⊔ b = unit := rfl -@[simp, nolint simpNF] +@[simp] theorem inf_eq : a ⊓ b = unit := rfl @@ -1096,16 +1096,15 @@ theorem inf_eq : a ⊓ b = unit := theorem compl_eq : aᶜ = unit := rfl -@[simp, nolint simpNF] +@[simp] theorem sdiff_eq : a \ b = unit := rfl -@[simp, nolint simpNF] +@[simp] theorem hnot_eq : ¬a = unit := rfl --- eligible for `dsimp` -@[simp, nolint simpNF] +@[simp] theorem himp_eq : a ⇨ b = unit := rfl diff --git a/Mathlib/Order/Interval/Basic.lean b/Mathlib/Order/Interval/Basic.lean index 014a33ce871f3..55171625e65d7 100644 --- a/Mathlib/Order/Interval/Basic.lean +++ b/Mathlib/Order/Interval/Basic.lean @@ -109,7 +109,6 @@ theorem mem_mk {hx : x.1 ≤ x.2} : a ∈ mk x hx ↔ x.1 ≤ a ∧ a ≤ x.2 := theorem mem_def : a ∈ s ↔ s.fst ≤ a ∧ a ≤ s.snd := Iff.rfl --- @[simp] -- Porting note: not in simpNF theorem coe_nonempty (s : NonemptyInterval α) : (s : Set α).Nonempty := nonempty_Icc.2 s.fst_le_snd diff --git a/Mathlib/Order/LiminfLimsup.lean b/Mathlib/Order/LiminfLimsup.lean index 25404163646dc..2576f34d60217 100644 --- a/Mathlib/Order/LiminfLimsup.lean +++ b/Mathlib/Order/LiminfLimsup.lean @@ -829,15 +829,13 @@ theorem HasBasis.limsup_eq_sInf_univ_of_empty {f : ι → α} {v : Filter ι} limsup f v = sInf univ := HasBasis.liminf_eq_sSup_univ_of_empty (α := αᵒᵈ) hv i hi h'i --- Porting note: simp_nf linter incorrectly says: lhs does not simplify when using simp on itself. -@[simp, nolint simpNF] +@[simp] theorem liminf_nat_add (f : ℕ → α) (k : ℕ) : liminf (fun i => f (i + k)) atTop = liminf f atTop := by change liminf (f ∘ (· + k)) atTop = liminf f atTop rw [liminf, liminf, ← map_map, map_add_atTop_eq_nat] --- Porting note: simp_nf linter incorrectly says: lhs does not simplify when using simp on itself. -@[simp, nolint simpNF] +@[simp] theorem limsup_nat_add (f : ℕ → α) (k : ℕ) : limsup (fun i => f (i + k)) atTop = limsup f atTop := @liminf_nat_add αᵒᵈ _ f k diff --git a/Mathlib/Order/SupIndep.lean b/Mathlib/Order/SupIndep.lean index 7e9d4b4cdbb85..68e516ec2adec 100644 --- a/Mathlib/Order/SupIndep.lean +++ b/Mathlib/Order/SupIndep.lean @@ -180,19 +180,7 @@ theorem SupIndep.attach (hs : s.SupIndep f) : s.attach.SupIndep fun a => f a := obtain ⟨j, hj, hji⟩ := hi' rwa [Subtype.ext hji] at hj -/- -Porting note: simpNF linter returns - -"Left-hand side does not simplify, when using the simp lemma on itself." - -However, simp does indeed solve the following. - -example {α ι} [Lattice α] [OrderBot α] (s : Finset ι) (f : ι → α) : - (s.attach.SupIndep fun a => f a) ↔ s.SupIndep f := by simp - -See https://github.com/leanprover-community/batteries/issues/71 for the simpNF issue. --/ -@[simp, nolint simpNF] +@[simp] theorem supIndep_attach : (s.attach.SupIndep fun a => f a) ↔ s.SupIndep f := by refine ⟨fun h t ht i his hit => ?_, SupIndep.attach⟩ classical diff --git a/Mathlib/RingTheory/DedekindDomain/SInteger.lean b/Mathlib/RingTheory/DedekindDomain/SInteger.lean index 83b4fe6cf9021..7568c704d46d4 100644 --- a/Mathlib/RingTheory/DedekindDomain/SInteger.lean +++ b/Mathlib/RingTheory/DedekindDomain/SInteger.lean @@ -96,7 +96,6 @@ theorem unit_valuation_eq_one (x : S.unit K) {v : HeightOneSpectrum R} (hv : v v.valuation ((x : Kˣ) : K) = 1 := x.property v hv --- Porting note: `apply_inv_coe` fails the simpNF linter /-- The group of `S`-units is the group of units of the ring of `S`-integers. -/ @[simps apply_val_coe symm_apply_coe] def unitEquivUnitsInteger : S.unit K ≃* (S.integer K)ˣ where diff --git a/Mathlib/RingTheory/Ideal/Operations.lean b/Mathlib/RingTheory/Ideal/Operations.lean index 13dfba254995d..7218ec0a9688e 100644 --- a/Mathlib/RingTheory/Ideal/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Operations.lean @@ -287,9 +287,7 @@ variable {R : Type u} [Semiring R] theorem add_eq_sup {I J : Ideal R} : I + J = I ⊔ J := rfl --- dsimp loops when applying this lemma to its LHS, --- probably https://github.com/leanprover/lean4/pull/2867 -@[simp, nolint simpNF] +@[simp] theorem zero_eq_bot : (0 : Ideal R) = ⊥ := rfl diff --git a/Mathlib/RingTheory/Valuation/Basic.lean b/Mathlib/RingTheory/Valuation/Basic.lean index e4f462f138500..fe79b5797d6c9 100644 --- a/Mathlib/RingTheory/Valuation/Basic.lean +++ b/Mathlib/RingTheory/Valuation/Basic.lean @@ -128,7 +128,7 @@ theorem coe_mk (f : R →*₀ Γ₀) (h) : ⇑(Valuation.mk f h) = f := rfl theorem toFun_eq_coe (v : Valuation R Γ₀) : v.toFun = v := rfl -@[simp] -- Porting note: requested by simpNF as toFun_eq_coe LHS simplifies +@[simp] theorem toMonoidWithZeroHom_coe_eq_coe (v : Valuation R Γ₀) : (v.toMonoidWithZeroHom : R → Γ₀) = v := rfl diff --git a/Mathlib/RingTheory/Valuation/ValuationSubring.lean b/Mathlib/RingTheory/Valuation/ValuationSubring.lean index 867c268d02ac5..813f07bdbde31 100644 --- a/Mathlib/RingTheory/Valuation/ValuationSubring.lean +++ b/Mathlib/RingTheory/Valuation/ValuationSubring.lean @@ -50,9 +50,6 @@ instance : SetLike (ValuationSubring K) K where replace h := SetLike.coe_injective' h congr --- Porting note https://github.com/leanprover-community/mathlib4/issues/10959 --- simp cannot prove this -@[simp, nolint simpNF] theorem mem_carrier (x : K) : x ∈ A.carrier ↔ x ∈ A := Iff.refl _ @[simp] @@ -668,9 +665,9 @@ def unitsModPrincipalUnitsEquivResidueFieldUnits : local instance : MulOneClass ({ x // x ∈ unitGroup A } ⧸ Subgroup.comap (Subgroup.subtype (unitGroup A)) (principalUnitGroup A)) := inferInstance --- @[simp] -- Porting note: not in simpNF theorem unitsModPrincipalUnitsEquivResidueFieldUnits_comp_quotientGroup_mk : - A.unitsModPrincipalUnitsEquivResidueFieldUnits.toMonoidHom.comp (QuotientGroup.mk' _) = + (A.unitsModPrincipalUnitsEquivResidueFieldUnits : _ ⧸ Subgroup.comap _ _ →* _).comp + (QuotientGroup.mk' (A.principalUnitGroup.subgroupOf A.unitGroup)) = A.unitGroupToResidueFieldUnits := rfl theorem unitsModPrincipalUnitsEquivResidueFieldUnits_comp_quotientGroup_mk_apply @@ -793,7 +790,6 @@ namespace Valuation variable {Γ : Type*} [LinearOrderedCommGroupWithZero Γ] (v : Valuation K Γ) (x : Kˣ) --- @[simp] -- Porting note: not in simpNF theorem mem_unitGroup_iff : x ∈ v.valuationSubring.unitGroup ↔ v x = 1 := IsEquiv.eq_one_iff_eq_one (Valuation.isEquiv_valuation_valuationSubring _).symm From 7e3651501ccaac21324762fb104e0f12b7143559 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 23 Nov 2024 23:31:35 +0000 Subject: [PATCH 254/829] feat: more `bound` lemmas (#19285) From LeanAPAP --- Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean index bbd7e7951179d..488d67de45348 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean @@ -983,6 +983,11 @@ lemma mul_lt_one_of_nonneg_of_lt_one_right [MulPosMono M₀] (ha : a ≤ 1) (hb section variable [ZeroLEOneClass M₀] [PosMulMono M₀] [MulPosMono M₀] +@[bound] +protected lemma Bound.one_lt_mul : 1 ≤ a ∧ 1 < b ∨ 1 < a ∧ 1 ≤ b → 1 < a * b := by + rintro (⟨ha, hb⟩ | ⟨ha, hb⟩); exacts [one_lt_mul ha hb, one_lt_mul_of_lt_of_le ha hb] + +@[bound] lemma mul_le_one₀ (ha : a ≤ 1) (hb₀ : 0 ≤ b) (hb : b ≤ 1) : a * b ≤ 1 := one_mul (1 : M₀) ▸ mul_le_mul ha hb hb₀ zero_le_one @@ -1358,6 +1363,8 @@ lemma inv_le_iff_one_le_mul₀' (ha : 0 < a) : a⁻¹ ≤ b ↔ 1 ≤ a * b := b lemma one_le_inv₀ (ha : 0 < a) : 1 ≤ a⁻¹ ↔ a ≤ 1 := by simpa using one_le_inv_mul₀ ha (b := 1) lemma inv_le_one₀ (ha : 0 < a) : a⁻¹ ≤ 1 ↔ 1 ≤ a := by simpa using inv_mul_le_one₀ ha (b := 1) +@[bound] alias ⟨_, Bound.one_le_inv₀⟩ := one_le_inv₀ + @[bound] lemma inv_le_one_of_one_le₀ (ha : 1 ≤ a) : a⁻¹ ≤ 1 := (inv_le_one₀ <| zero_lt_one.trans_le ha).2 ha From 13f8b501146167f8c599e41a7f9b49054a16f6a7 Mon Sep 17 00:00:00 2001 From: leanprover-community-bot-assistant Date: Sun, 24 Nov 2024 00:16:54 +0000 Subject: [PATCH 255/829] chore(scripts): update nolints.json (#19428) I am happy to remove some nolints for you! --- scripts/nolints.json | 13 ------------- 1 file changed, 13 deletions(-) diff --git a/scripts/nolints.json b/scripts/nolints.json index a64eacb27c056..a698a174e1354 100644 --- a/scripts/nolints.json +++ b/scripts/nolints.json @@ -575,20 +575,7 @@ ["docBlame", "CategoryTheory.Limits.MultispanIndex.right"], ["docBlame", "CategoryTheory.Limits.MultispanIndex.snd"], ["docBlame", "CategoryTheory.Limits.MultispanIndex.sndFrom"], - ["docBlame", "CategoryTheory.Limits.PreservesColimit.preserves"], - ["docBlame", "CategoryTheory.Limits.PreservesLimit.preserves"], - ["docBlame", "CategoryTheory.Limits.ReflectsColimit.reflects"], - ["docBlame", "CategoryTheory.Limits.ReflectsColimitsOfShape.reflectsColimit"], - ["docBlame", - "CategoryTheory.Limits.ReflectsColimitsOfSize.reflectsColimitsOfShape"], - ["docBlame", "CategoryTheory.Limits.ReflectsLimit.reflects"], - ["docBlame", "CategoryTheory.Limits.ReflectsLimitsOfShape.reflectsLimit"], - ["docBlame", - "CategoryTheory.Limits.ReflectsLimitsOfSize.reflectsLimitsOfShape"], ["docBlame", "CategoryTheory.Monad.CreatesColimitOfIsSplitPair.out"], - ["docBlame", "CategoryTheory.Monad.PreservesColimitOfIsReflexivePair.out"], - ["docBlame", "CategoryTheory.Monad.PreservesColimitOfIsSplitPair.out"], - ["docBlame", "CategoryTheory.Monad.ReflectsColimitOfIsSplitPair.out"], ["docBlame", "CategoryTheory.Pretriangulated.Triangle.homMk"], ["docBlame", "CategoryTheory.Pretriangulated.Triangle.isoMk"], ["docBlame", "CategoryTheory.Triangulated.Octahedron.m₁"], From 3ece930d0a4a55679efa52b1a825ac93b2469a06 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sun, 24 Nov 2024 07:24:05 +0000 Subject: [PATCH 256/829] chore: import Plausible in Tactic.Common (#19429) --- Mathlib/Tactic/Common.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/Tactic/Common.lean b/Mathlib/Tactic/Common.lean index 90340cd10a34e..3c61c4a884912 100644 --- a/Mathlib/Tactic/Common.lean +++ b/Mathlib/Tactic/Common.lean @@ -4,9 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ --- First import Aesop and Qq +-- First import Aesop, Qq, and Plausible import Aesop import Qq +import Plausible -- Tools for analysing imports, like `#find_home`, `#minimize_imports`, ... import ImportGraph.Imports From 83ce5b0136c9d9b856df5fc8fb4baa1c0527372f Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Sun, 24 Nov 2024 10:18:35 +0000 Subject: [PATCH 257/829] feat(CategoryTheory/Over): more API for `post` (#19312) --- Mathlib/CategoryTheory/Comma/Over.lean | 181 +++++++++++++++++++++++++ 1 file changed, 181 insertions(+) diff --git a/Mathlib/CategoryTheory/Comma/Over.lean b/Mathlib/CategoryTheory/Comma/Over.lean index e37bf7ec99866..305664eb69cad 100644 --- a/Mathlib/CategoryTheory/Comma/Over.lean +++ b/Mathlib/CategoryTheory/Comma/Over.lean @@ -161,6 +161,13 @@ theorem map_map_left : ((map f).map g).left = g.left := rfl end +/-- If `f` is an isomorphism, `map f` is an equivalence of categories. -/ +def mapIso {Y : T} (f : X ≅ Y) : Over X ≌ Over Y := + Comma.mapRightIso _ <| Discrete.natIso fun _ ↦ f + +@[simp] lemma mapIso_functor {Y : T} (f : X ≅ Y) : (mapIso f).functor = map f.hom := rfl +@[simp] lemma mapIso_inverse {Y : T} (f : X ≅ Y) : (mapIso f).inverse = map f.inv := rfl + section coherences /-! This section proves various equalities between functors that @@ -186,6 +193,7 @@ theorem mapId_eq (Y : T) : map (𝟙 Y) = 𝟭 _ := by simp /-- The natural isomorphism arising from `mapForget_eq`. -/ +@[simps!] def mapId (Y : T) : map (𝟙 Y) ≅ 𝟭 _ := eqToIso (mapId_eq Y) -- NatIso.ofComponents fun X => isoMk (Iso.refl _) @@ -215,9 +223,16 @@ theorem mapComp_eq {X Y Z : T} (f : X ⟶ Y) (g : Y ⟶ Z) : simp /-- The natural isomorphism arising from `mapComp_eq`. -/ +@[simps!] def mapComp {X Y Z : T} (f : X ⟶ Y) (g : Y ⟶ Z) : map (f ≫ g) ≅ (map f) ⋙ (map g) := eqToIso (mapComp_eq f g) +/-- If `f = g`, then `map f` is naturally isomorphic to `map g`. -/ +@[simps!] +def mapCongr {X Y : T} (f g : X ⟶ Y) (h : f = g) : + map f ≅ map g := + NatIso.ofComponents (fun A ↦ eqToIso (by rw [h])) + variable (T) in /-- The functor defined by the over categories.-/ @[simps] def mapFunctor : T ⥤ Cat where @@ -317,6 +332,57 @@ def post (F : T ⥤ D) : Over X ⥤ Over (F.obj X) where map f := Over.homMk (F.map f.left) (by simp only [Functor.id_obj, mk_left, Functor.const_obj_obj, mk_hom, ← F.map_comp, w]) +lemma post_comp {E : Type*} [Category E] (F : T ⥤ D) (G : D ⥤ E) : + post (X := X) (F ⋙ G) = post (X := X) F ⋙ post G := + rfl + +/-- `post (F ⋙ G)` is isomorphic (actually equal) to `post F ⋙ post G`. -/ +@[simps!] +def postComp {E : Type*} [Category E] (F : T ⥤ D) (G : D ⥤ E) : + post (X := X) (F ⋙ G) ≅ post F ⋙ post G := + NatIso.ofComponents (fun X ↦ Iso.refl _) + +/-- A natural transformation `F ⟶ G` induces a natural transformation on +`Over X` up to `Under.map`. -/ +@[simps] +def postMap {F G : T ⥤ D} (e : F ⟶ G) : post F ⋙ map (e.app X) ⟶ post G where + app Y := Over.homMk (e.app Y.left) + +/-- If `F` and `G` are naturally isomorphic, then `Over.post F` and `Over.post G` are also naturally +isomorphic up to `Over.map` -/ +@[simps!] +def postCongr {F G : T ⥤ D} (e : F ≅ G) : post F ⋙ map (e.hom.app X) ≅ post G := + NatIso.ofComponents (fun A ↦ Over.isoMk (e.app A.left)) + +variable (X) (F : T ⥤ D) + +instance [F.Faithful] : (Over.post (X := X) F).Faithful where + map_injective {A B} f g h := by + ext + exact F.map_injective (congrArg CommaMorphism.left h) + +instance [F.Faithful] [F.Full] : (Over.post (X := X) F).Full where + map_surjective {A B} f := by + obtain ⟨a, ha⟩ := F.map_surjective f.left + have w : a ≫ B.hom = A.hom := F.map_injective <| by simpa [ha] using Over.w _ + exact ⟨Over.homMk a, by ext; simpa⟩ + +instance [F.Full] [F.EssSurj] : (Over.post (X := X) F).EssSurj where + mem_essImage B := by + obtain ⟨A', ⟨e⟩⟩ := Functor.EssSurj.mem_essImage (F := F) B.left + obtain ⟨f, hf⟩ := F.map_surjective (e.hom ≫ B.hom) + exact ⟨Over.mk f, ⟨Over.isoMk e⟩⟩ + +instance [F.IsEquivalence] : (Over.post (X := X) F).IsEquivalence where + +/-- An equivalence of categories induces an equivalence on over categories. -/ +@[simps] +def postEquiv (F : T ≌ D) : Over X ≌ Over (F.functor.obj X) where + functor := Over.post F.functor + inverse := Over.post (X := F.functor.obj X) F.inverse ⋙ Over.map (F.unitIso.inv.app X) + unitIso := NatIso.ofComponents (fun A ↦ Over.isoMk (F.unitIso.app A.left)) + counitIso := NatIso.ofComponents (fun A ↦ Over.isoMk (F.counitIso.app A.left)) + end Over namespace CostructuredArrow @@ -457,6 +523,13 @@ theorem map_map_right : ((map f).map g).right = g.right := rfl end +/-- If `f` is an isomorphism, `map f` is an equivalence of categories. -/ +def mapIso {Y : T} (f : X ≅ Y) : Under Y ≌ Under X := + Comma.mapLeftIso _ <| Discrete.natIso fun _ ↦ f.symm + +@[simp] lemma mapIso_functor {Y : T} (f : X ≅ Y) : (mapIso f).functor = map f.hom := rfl +@[simp] lemma mapIso_inverse {Y : T} (f : X ≅ Y) : (mapIso f).inverse = map f.inv := rfl + section coherences /-! This section proves various equalities between functors that @@ -476,6 +549,7 @@ theorem mapId_eq (Y : T) : map (𝟙 Y) = 𝟭 _ := by simp /-- Mapping by the identity morphism is just the identity functor. -/ +@[simps!] def mapId (Y : T) : map (𝟙 Y) ≅ 𝟭 _ := eqToIso (mapId_eq Y) /-- Mapping by `f` and then forgetting is the same as forgetting. -/ @@ -504,9 +578,16 @@ theorem mapComp_eq {X Y Z : T} (f : X ⟶ Y) (g : Y ⟶ Z) : simp /-- The natural isomorphism arising from `mapComp_eq`. -/ +@[simps!] def mapComp {Y Z : T} (f : X ⟶ Y) (g : Y ⟶ Z) : map (f ≫ g) ≅ map g ⋙ map f := eqToIso (mapComp_eq f g) +/-- If `f = g`, then `map f` is naturally isomorphic to `map g`. -/ +@[simps!] +def mapCongr {X Y : T} (f g : X ⟶ Y) (h : f = g) : + map f ≅ map g := + NatIso.ofComponents (fun A ↦ eqToIso (by rw [h])) + variable (T) in /-- The functor defined by the under categories.-/ @[simps] def mapFunctor : Tᵒᵖ ⥤ Cat where @@ -570,6 +651,58 @@ def post {X : T} (F : T ⥤ D) : Under X ⥤ Under (F.obj X) where map f := Under.homMk (F.map f.right) (by simp only [Functor.id_obj, Functor.const_obj_obj, mk_right, mk_hom, ← F.map_comp, w]) +lemma post_comp {E : Type*} [Category E] (F : T ⥤ D) (G : D ⥤ E) : + post (X := X) (F ⋙ G) = post (X := X) F ⋙ post G := + rfl + +/-- `post (F ⋙ G)` is isomorphic (actually equal) to `post F ⋙ post G`. -/ +@[simps!] +def postComp {E : Type*} [Category E] (F : T ⥤ D) (G : D ⥤ E) : + post (X := X) (F ⋙ G) ≅ post F ⋙ post G := + NatIso.ofComponents (fun X ↦ Iso.refl _) + +/-- A natural transformation `F ⟶ G` induces a natural transformation on +`Under X` up to `Under.map`. -/ +@[simps] +def postMap {F G : T ⥤ D} (e : F ⟶ G) : post (X := X) F ⟶ post G ⋙ map (e.app X) where + app Y := Under.homMk (e.app Y.right) + +/-- If `F` and `G` are naturally isomorphic, then `Under.post F` and `Under.post G` are also +naturally isomorphic up to `Under.map` -/ +@[simps!] +def postCongr {F G : T ⥤ D} (e : F ≅ G) : post F ≅ post G ⋙ map (e.hom.app X) := + NatIso.ofComponents (fun A ↦ Under.isoMk (e.app A.right)) + +variable (X) (F : T ⥤ D) + +instance [F.Faithful] : (Under.post (X := X) F).Faithful where + map_injective {A B} f g h := by + ext + exact F.map_injective (congrArg CommaMorphism.right h) + +instance [F.Faithful] [F.Full] : (Under.post (X := X) F).Full where + map_surjective {A B} f := by + obtain ⟨a, ha⟩ := F.map_surjective f.right + dsimp at a + have w : A.hom ≫ a = B.hom := F.map_injective <| by simpa [ha] using Under.w f + exact ⟨Under.homMk a, by ext; simpa⟩ + +instance [F.Full] [F.EssSurj] : (Under.post (X := X) F).EssSurj where + mem_essImage B := by + obtain ⟨B', ⟨e⟩⟩ := Functor.EssSurj.mem_essImage (F := F) B.right + obtain ⟨f, hf⟩ := F.map_surjective (B.hom ≫ e.inv) + exact ⟨Under.mk f, ⟨Under.isoMk e⟩⟩ + +instance [F.IsEquivalence] : (Under.post (X := X) F).IsEquivalence where + +/-- An equivalence of categories induces an equivalence on under categories. -/ +@[simps] +def postEquiv (F : T ≌ D) : Under X ≌ Under (F.functor.obj X) where + functor := post F.functor + inverse := post (X := F.functor.obj X) F.inverse ⋙ Under.map (F.unitIso.hom.app X) + unitIso := NatIso.ofComponents (fun A ↦ Under.isoMk (F.unitIso.app A.right)) + counitIso := NatIso.ofComponents (fun A ↦ Under.isoMk (F.counitIso.app A.right)) + end Under namespace StructuredArrow @@ -774,4 +907,52 @@ def ofDiagEquivalence' (X : T × T) : end CostructuredArrow +section Opposite + +open Opposite + +variable (X : T) + +/-- The canonical functor by reversing structure arrows. -/ +@[simps] +def Over.opToOpUnder : Over (op X) ⥤ (Under X)ᵒᵖ where + obj Y := ⟨Under.mk Y.hom.unop⟩ + map {Z Y} f := ⟨Under.homMk (f.left.unop) (by dsimp; rw [← unop_comp, Over.w])⟩ + +/-- The canonical functor by reversing structure arrows. -/ +@[simps] +def Under.opToOverOp : (Under X)ᵒᵖ ⥤ Over (op X) where + obj Y := Over.mk (Y.unop.hom.op) + map {Z Y} f := Over.homMk f.unop.right.op <| by dsimp; rw [← Under.w f.unop, op_comp] + +/-- `Over.opToOpUnder` is an equivalence of categories. -/ +@[simps] +def Over.opEquivOpUnder : Over (op X) ≌ (Under X)ᵒᵖ where + functor := Over.opToOpUnder X + inverse := Under.opToOverOp X + unitIso := Iso.refl _ + counitIso := Iso.refl _ + +/-- The canonical functor by reversing structure arrows. -/ +@[simps] +def Under.opToOpOver : Under (op X) ⥤ (Over X)ᵒᵖ where + obj Y := ⟨Over.mk Y.hom.unop⟩ + map {Z Y} f := ⟨Over.homMk (f.right.unop) (by dsimp; rw [← unop_comp, Under.w])⟩ + +/-- The canonical functor by reversing structure arrows. -/ +@[simps] +def Over.opToUnderOp : (Over X)ᵒᵖ ⥤ Under (op X) where + obj Y := Under.mk (Y.unop.hom.op) + map {Z Y} f := Under.homMk f.unop.left.op <| by dsimp; rw [← Over.w f.unop, op_comp] + +/-- `Under.opToOpOver` is an equivalence of categories. -/ +@[simps] +def Under.opEquivOpOver : Under (op X) ≌ (Over X)ᵒᵖ where + functor := Under.opToOpOver X + inverse := Over.opToUnderOp X + unitIso := Iso.refl _ + counitIso := Iso.refl _ + +end Opposite + end CategoryTheory From 6c225037eb50056d12f06584a2dd4ee0008910d9 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Sun, 24 Nov 2024 13:54:52 +0000 Subject: [PATCH 258/829] chore(AlgebraicGeometry) add `Scheme.Hom.appTop` (#19130) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit to hide the defeq abuse `f ⁻¹ᵁ ⊤ = ⊤` --- Mathlib/AlgebraicGeometry/AffineScheme.lean | 52 +++---- Mathlib/AlgebraicGeometry/AffineSpace.lean | 135 ++++++++---------- .../GammaSpecAdjunction.lean | 25 ++-- .../AlgebraicGeometry/Morphisms/Affine.lean | 2 +- .../Morphisms/AffineAnd.lean | 29 ++-- .../Morphisms/ClosedImmersion.lean | 28 ++-- .../AlgebraicGeometry/Morphisms/Finite.lean | 2 +- .../AlgebraicGeometry/Morphisms/IsIso.lean | 4 +- .../Morphisms/RingHomProperties.lean | 38 ++--- .../Morphisms/SurjectiveOnStalks.lean | 1 - Mathlib/AlgebraicGeometry/Restrict.lean | 25 +++- Mathlib/AlgebraicGeometry/Scheme.lean | 33 ++++- Mathlib/AlgebraicGeometry/Stalk.lean | 9 +- .../AlgebraicGeometry/ValuativeCriterion.lean | 6 +- 14 files changed, 214 insertions(+), 175 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index a83addf8247e5..96002facc9d47 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -67,12 +67,12 @@ def Scheme.isoSpec (X : Scheme) [IsAffine X] : X ≅ Spec Γ(X, ⊤) := @[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 + X.isoSpec.hom ≫ Spec.map (f.appTop) = f ≫ Y.isoSpec.hom := by 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 + Spec.map (f.appTop) ≫ Y.isoSpec.inv = X.isoSpec.inv ≫ f := by rw [Iso.eq_inv_comp, isoSpec, asIso_hom, ← Scheme.toSpecΓ_naturality_assoc, isoSpec, asIso_inv, IsIso.hom_inv_id, Category.comp_id] @@ -112,13 +112,13 @@ theorem isAffine_of_isIso {X Y : Scheme} (f : X ⟶ Y) [IsIso f] [h : IsAffine Y to the arrow of the morphism on prime spectra induced by the map on global sections. -/ noncomputable def arrowIsoSpecΓOfIsAffine {X Y : Scheme} [IsAffine X] [IsAffine Y] (f : X ⟶ Y) : - Arrow.mk f ≅ Arrow.mk (Spec.map (Scheme.Γ.map f.op)) := + Arrow.mk f ≅ Arrow.mk (Spec.map f.appTop) := Arrow.isoMk X.isoSpec Y.isoSpec (ΓSpec.adjunction.unit_naturality _) /-- If `f : A ⟶ B` is a ring homomorphism, the corresponding arrow is isomorphic to the arrow of the morphism induced on global sections by the map on prime spectra. -/ def arrowIsoΓSpecOfIsAffine {A B : CommRingCat} (f : A ⟶ B) : - Arrow.mk f ≅ Arrow.mk ((Spec.map f).app ⊤) := + Arrow.mk f ≅ Arrow.mk ((Spec.map f).appTop) := Arrow.isoMk (Scheme.ΓSpecIso _).symm (Scheme.ΓSpecIso _).symm (Scheme.ΓSpecIso_inv_naturality f).symm @@ -134,7 +134,7 @@ theorem Scheme.isoSpec_Spec (R : CommRingCat.{u}) : (Spec R).isoSpec.inv = Spec.map (Scheme.ΓSpecIso R).inv := congr($(isoSpec_Spec R).inv) -lemma ext_of_isAffine {X Y : Scheme} [IsAffine Y] {f g : X ⟶ Y} (e : f.app ⊤ = g.app ⊤) : +lemma ext_of_isAffine {X Y : Scheme} [IsAffine Y] {f g : X ⟶ Y} (e : f.appTop = g.appTop) : f = g := by rw [← cancel_mono Y.toSpecΓ, Scheme.toSpecΓ_naturality, Scheme.toSpecΓ_naturality, e] @@ -315,21 +315,24 @@ lemma isoSpec_hom_base_apply (x : U) : congr 1 exact IsLocalRing.comap_closedPoint (U.stalkIso x).inv -lemma isoSpec_inv_app_top : - hU.isoSpec.inv.app ⊤ = U.topIso.hom ≫ (Scheme.ΓSpecIso Γ(X, U)).inv := by +lemma isoSpec_inv_appTop : + hU.isoSpec.inv.appTop = U.topIso.hom ≫ (Scheme.ΓSpecIso Γ(X, U)).inv := by simp only [Scheme.Opens.toScheme_presheaf_obj, isoSpec_inv, Scheme.isoSpec, asIso_inv, - Scheme.comp_coeBase, Opens.map_comp_obj, Opens.map_top, Scheme.comp_app, Scheme.inv_app_top, + Scheme.comp_coeBase, Opens.map_comp_obj, Opens.map_top, Scheme.comp_app, Scheme.inv_appTop, Scheme.Opens.topIso_hom, Scheme.ΓSpecIso_inv_naturality, IsIso.inv_comp_eq] - rw [Scheme.toSpecΓ_app_top] + rw [Scheme.toSpecΓ_appTop] erw [Iso.hom_inv_id_assoc] -lemma isoSpec_hom_app_top : - hU.isoSpec.hom.app ⊤ = (Scheme.ΓSpecIso Γ(X, U)).hom ≫ U.topIso.inv := by - have := congr(inv $hU.isoSpec_inv_app_top) +lemma isoSpec_hom_appTop : + hU.isoSpec.hom.appTop = (Scheme.ΓSpecIso Γ(X, U)).hom ≫ U.topIso.inv := by + have := congr(inv $hU.isoSpec_inv_appTop) rw [IsIso.inv_comp, IsIso.Iso.inv_inv, IsIso.Iso.inv_hom] at this have := (Scheme.Γ.map_inv hU.isoSpec.inv.op).trans this rwa [← op_inv, IsIso.Iso.inv_inv] at this +@[deprecated (since := "2024-11-16")] alias isoSpec_inv_app_top := isoSpec_inv_appTop +@[deprecated (since := "2024-11-16")] alias isoSpec_hom_app_top := isoSpec_hom_appTop + /-- The open immersion `Spec Γ(X, U) ⟶ X` for an affine `U`. -/ def fromSpec : Spec Γ(X, U) ⟶ X := @@ -364,7 +367,7 @@ theorem map_fromSpec {V : X.Opens} (hV : IsAffineOpen V) (f : op U ⟶ op V) : conv_rhs => rw [fromSpec, ← X.homOfLE_ι (V := U) f.unop.le, isoSpec_inv, Category.assoc, ← Scheme.isoSpec_inv_naturality_assoc, - ← Spec.map_comp_assoc, Scheme.homOfLE_app, ← Functor.map_comp] + ← Spec.map_comp_assoc, Scheme.homOfLE_appTop, ← Functor.map_comp] rw [fromSpec, isoSpec_inv, Category.assoc, ← Spec.map_comp_assoc, ← Functor.map_comp] rfl @@ -377,21 +380,22 @@ lemma Spec_map_appLE_fromSpec (f : X ⟶ Y) {V : X.Opens} {U : Y.Opens} simp_rw [← Scheme.homOfLE_ι _ i] rw [Category.assoc, ← morphismRestrict_ι, ← Category.assoc _ (f ∣_ U) U.ι, ← @Scheme.isoSpec_inv_naturality_assoc, - ← Spec.map_comp_assoc, ← Spec.map_comp_assoc, Scheme.comp_app, morphismRestrict_app, - Scheme.homOfLE_app, Scheme.Hom.app_eq_appLE, Scheme.Hom.appLE_map, + ← Spec.map_comp_assoc, ← Spec.map_comp_assoc, Scheme.comp_appTop, morphismRestrict_appTop, + Scheme.homOfLE_appTop, Scheme.Hom.app_eq_appLE, Scheme.Hom.appLE_map, Scheme.Hom.appLE_map, Scheme.Hom.appLE_map, Scheme.Hom.map_appLE] lemma fromSpec_top [IsAffine X] : (isAffineOpen_top X).fromSpec = X.isoSpec.inv := by - rw [fromSpec, isoSpec_inv, Category.assoc, ← @Scheme.isoSpec_inv_naturality, Scheme.Opens.ι_app, - ← Spec.map_comp_assoc, ← X.presheaf.map_comp, ← op_comp, eqToHom_comp_homOfLE, - ← eqToHom_eq_homOfLE rfl, eqToHom_refl, op_id, X.presheaf.map_id, Spec.map_id, Category.id_comp] + rw [fromSpec, isoSpec_inv, Category.assoc, ← @Scheme.isoSpec_inv_naturality, + ← Spec.map_comp_assoc, Scheme.Opens.ι_appTop, ← X.presheaf.map_comp, ← op_comp, + eqToHom_comp_homOfLE, ← eqToHom_eq_homOfLE rfl, eqToHom_refl, op_id, X.presheaf.map_id, + Spec.map_id, Category.id_comp] lemma fromSpec_app_of_le (V : X.Opens) (h : U ≤ V) : hU.fromSpec.app V = X.presheaf.map (homOfLE h).op ≫ (Scheme.ΓSpecIso Γ(X, U)).inv ≫ (Spec _).presheaf.map (homOfLE le_top).op := by have : U.ι ⁻¹ᵁ V = ⊤ := eq_top_iff.mpr fun x _ ↦ h x.2 rw [IsAffineOpen.fromSpec, Scheme.comp_app, Scheme.Opens.ι_app, Scheme.app_eq _ this, - IsAffineOpen.isoSpec_inv_app_top] + ← Scheme.Hom.appTop, IsAffineOpen.isoSpec_inv_appTop] simp only [Scheme.Opens.toScheme_presheaf_map, Scheme.Opens.topIso_hom, Category.assoc, ← X.presheaf.map_comp_assoc] rfl @@ -470,7 +474,7 @@ theorem ΓSpecIso_hom_fromSpec_app : simp only [fromSpec, Scheme.comp_coeBase, Opens.map_comp_obj, Scheme.comp_app, Scheme.Opens.ι_app_self, eqToHom_op, Scheme.app_eq _ U.ι_preimage_self, Scheme.Opens.toScheme_presheaf_map, eqToHom_unop, eqToHom_map U.ι.opensFunctor, Opens.map_top, - isoSpec_inv_app_top, Scheme.Opens.topIso_hom, Category.assoc, ← Functor.map_comp_assoc, + isoSpec_inv_appTop, Scheme.Opens.topIso_hom, Category.assoc, ← Functor.map_comp_assoc, eqToHom_trans, eqToHom_refl, X.presheaf.map_id, Category.id_comp, Iso.hom_inv_id_assoc] @[elementwise] @@ -953,10 +957,10 @@ lemma specTargetImageRingHom_surjective : Function.Surjective (specTargetImageRi Ideal.Quotient.mk_surjective lemma specTargetImageFactorization_app_injective : - Function.Injective <| (specTargetImageFactorization f).app ⊤ := by + Function.Injective <| (specTargetImageFactorization f).appTop := by let φ : A ⟶ Γ(X, ⊤) := (((ΓSpec.adjunction).homEquiv X (op A)).symm f).unop let φ' : specTargetImage f ⟶ Scheme.Γ.obj (op X) := RingHom.kerLift φ - show Function.Injective <| ((ΓSpec.adjunction.homEquiv X _) φ'.op).app ⊤ + show Function.Injective <| ((ΓSpec.adjunction.homEquiv X _) φ'.op).appTop rw [ΓSpec_adjunction_homEquiv_eq] apply (RingHom.kerLift_injective φ).comp exact ((ConcreteCategory.isIso_iff_bijective (Scheme.ΓSpecIso _).hom).mp inferInstance).injective @@ -993,7 +997,7 @@ def affineTargetImageInclusion (f : X ⟶ Y) : affineTargetImage f ⟶ Y := Spec.map (specTargetImageRingHom (f ≫ Y.isoSpec.hom)) ≫ Y.isoSpec.inv lemma affineTargetImageInclusion_app_surjective : - Function.Surjective <| (affineTargetImageInclusion f).app ⊤ := by + Function.Surjective <| (affineTargetImageInclusion f).appTop := by simp only [Scheme.comp_coeBase, Opens.map_comp_obj, Opens.map_top, Scheme.comp_app, CommRingCat.coe_comp, affineTargetImageInclusion] apply Function.Surjective.comp @@ -1013,7 +1017,7 @@ def affineTargetImageFactorization (f : X ⟶ Y) : X ⟶ affineTargetImage f := specTargetImageFactorization (f ≫ Y.isoSpec.hom) lemma affineTargetImageFactorization_app_injective : - Function.Injective <| (affineTargetImageFactorization f).app ⊤ := + Function.Injective <| (affineTargetImageFactorization f).appTop := specTargetImageFactorization_app_injective (f ≫ Y.isoSpec.hom) @[reassoc (attr := simp)] diff --git a/Mathlib/AlgebraicGeometry/AffineSpace.lean b/Mathlib/AlgebraicGeometry/AffineSpace.lean index 8e43ca277f77f..805e43c4928d2 100644 --- a/Mathlib/AlgebraicGeometry/AffineSpace.lean +++ b/Mathlib/AlgebraicGeometry/AffineSpace.lean @@ -71,7 +71,7 @@ Use `homOverEquiv` instead. -/ @[simps] def toSpecMvPolyIntEquiv : (X ⟶ Spec ℤ[n]) ≃ (n → Γ(X, ⊤)) where - toFun f i := f.app ⊤ ((Scheme.ΓSpecIso ℤ[n]).inv (.X i)) + toFun f i := f.appTop ((Scheme.ΓSpecIso ℤ[n]).inv (.X i)) invFun v := X.toSpecΓ ≫ Spec.map (MvPolynomial.eval₂Hom ((algebraMap ℤ _).comp ULift.ringEquiv.toRingHom) v) left_inv f := by @@ -80,7 +80,7 @@ def toSpecMvPolyIntEquiv : (X ⟶ Spec ℤ[n]) ≃ (n → Γ(X, ⊤)) where rw [Adjunction.homEquiv_symm_apply, Adjunction.homEquiv_symm_apply] simp only [Functor.rightOp_obj, Scheme.Γ_obj, Scheme.Spec_obj, algebraMap_int_eq, RingEquiv.toRingHom_eq_coe, TopologicalSpace.Opens.map_top, Functor.rightOp_map, op_comp, - Scheme.Γ_map, unop_comp, Quiver.Hom.unop_op, Scheme.comp_app, Scheme.toSpecΓ_app_top, + Scheme.Γ_map, unop_comp, Quiver.Hom.unop_op, Scheme.comp_app, Scheme.toSpecΓ_appTop, Scheme.ΓSpecIso_naturality, ΓSpec.adjunction_counit_app, Category.assoc, Iso.cancel_iso_inv_left, ← Iso.eq_inv_comp] apply of_mvPolynomial_int_ext @@ -91,13 +91,13 @@ def toSpecMvPolyIntEquiv : (X ⟶ Spec ℤ[n]) ≃ (n → Γ(X, ⊤)) where ext i simp only [algebraMap_int_eq, RingEquiv.toRingHom_eq_coe, Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, TopologicalSpace.Opens.map_top, Scheme.comp_app, - Scheme.toSpecΓ_app_top, Scheme.ΓSpecIso_naturality, CommRingCat.comp_apply, + Scheme.toSpecΓ_appTop, Scheme.ΓSpecIso_naturality, CommRingCat.comp_apply, CommRingCat.coe_of] erw [Iso.hom_inv_id_apply] rw [coe_eval₂Hom, eval₂_X] lemma toSpecMvPolyIntEquiv_comp {X Y : Scheme} (f : X ⟶ Y) (g : Y ⟶ Spec ℤ[n]) (i) : - toSpecMvPolyIntEquiv n (f ≫ g) i = f.app ⊤ (toSpecMvPolyIntEquiv n g i) := rfl + toSpecMvPolyIntEquiv n (f ≫ g) i = f.appTop (toSpecMvPolyIntEquiv n g i) := rfl variable {n} in /-- The standard coordinates of `𝔸(n; S)`. -/ @@ -122,16 +122,16 @@ lemma homOfVector_toSpecMvPoly : homOfVector f v ≫ toSpecMvPoly n S = (toSpecMvPolyIntEquiv n).symm v := pullback.lift_snd _ _ _ -@[local simp] -lemma homOfVector_app_top_coord (i) : - (homOfVector f v).app ⊤ (coord S i) = v i := by +@[simp] +lemma homOfVector_appTop_coord (i) : + (homOfVector f v).appTop (coord S i) = v i := by rw [coord, ← toSpecMvPolyIntEquiv_comp, homOfVector_toSpecMvPoly, Equiv.apply_symm_apply] @[ext 1100] lemma hom_ext {f g : X ⟶ 𝔸(n; S)} (h₁ : f ≫ 𝔸(n; S) ↘ S = g ≫ 𝔸(n; S) ↘ S) - (h₂ : ∀ i, f.app ⊤ (coord S i) = g.app ⊤ (coord S i)) : f = g := by + (h₂ : ∀ i, f.appTop (coord S i) = g.appTop (coord S i)) : f = g := by apply pullback.hom_ext h₁ show f ≫ toSpecMvPoly _ _ = g ≫ toSpecMvPoly _ _ apply (toSpecMvPolyIntEquiv n).injective @@ -141,8 +141,8 @@ lemma hom_ext {f g : X ⟶ 𝔸(n; S)} @[reassoc] lemma comp_homOfVector {X Y : Scheme} (v : n → Γ(Y, ⊤)) (f : X ⟶ Y) (g : Y ⟶ S) : - f ≫ homOfVector g v = homOfVector (f ≫ g) (f.app ⊤ ∘ v) := by - ext1 <;> simp [-TopologicalSpace.Opens.map_top]; rfl + f ≫ homOfVector g v = homOfVector (f ≫ g) (f.appTop ∘ v) := by + ext1 <;> simp end homOfVector @@ -155,13 +155,13 @@ instance (v : n → Γ(X, ⊤)) : (homOfVector (X ↘ S) v).IsOver S where /-- `S`-morphisms into `Spec 𝔸(n; S)` are equivalent to the choice of `n` global sections. -/ @[simps] def homOverEquiv : { f : X ⟶ 𝔸(n; S) // f.IsOver S } ≃ (n → Γ(X, ⊤)) where - toFun f i := f.1.app ⊤ (coord S i) + toFun f i := f.1.appTop (coord S i) invFun v := ⟨homOfVector (X ↘ S) v, inferInstance⟩ left_inv f := by ext : 2 · simp [f.2.1] - · rw [homOfVector_app_top_coord] - right_inv v := by ext i; simp [-TopologicalSpace.Opens.map_top, homOfVector_app_top_coord] + · rw [homOfVector_appTop_coord] + right_inv v := by ext i; simp [-TopologicalSpace.Opens.map_top, homOfVector_appTop_coord] variable (n) in /-- @@ -171,7 +171,7 @@ Also see `AffineSpace.SpecIso`. @[simps (config := .lemmasOnly) hom inv] def isoOfIsAffine [IsAffine S] : 𝔸(n; S) ≅ Spec (.of (MvPolynomial n Γ(S, ⊤))) where - hom := 𝔸(n; S).toSpecΓ ≫ Spec.map (eval₂Hom ((𝔸(n; S) ↘ S).app ⊤) (coord S)) + hom := 𝔸(n; S).toSpecΓ ≫ Spec.map (eval₂Hom ((𝔸(n; S) ↘ S).appTop) (coord S)) inv := homOfVector (Spec.map C ≫ S.isoSpec.inv) ((Scheme.ΓSpecIso (.of (MvPolynomial n Γ(S, ⊤)))).inv ∘ MvPolynomial.X) hom_inv_id := by @@ -180,48 +180,43 @@ def isoOfIsAffine [IsAffine S] : rw [← Spec.map_comp_assoc, CommRingCat.comp_eq_ring_hom_comp, eval₂Hom_comp_C, ← Scheme.toSpecΓ_naturality_assoc] simp [Scheme.isoSpec] - · simp only [Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, Scheme.comp_app, - CommRingCat.coe_comp_of, RingHom.coe_comp, Function.comp_apply, Scheme.id.base, - Scheme.id_app] - -- note: `rw [homOfVector_app_top_coord]` works but takes 3 more seconds - conv => enter [1, 2, 2]; rw [homOfVector_app_top_coord] - simp only [TopologicalSpace.Opens.map_top, Scheme.toSpecΓ_app_top, Function.comp_apply, + · simp only [Category.assoc, Scheme.comp_app, Scheme.comp_coeBase, + TopologicalSpace.Opens.map_comp_obj, TopologicalSpace.Opens.map_top, + Scheme.toSpecΓ_appTop, Scheme.ΓSpecIso_naturality, CommRingCat.comp_apply, + homOfVector_appTop_coord, Function.comp_apply, CommRingCat.coe_of, Scheme.id_app, CommRingCat.id_apply] - erw [elementwise_of% Scheme.ΓSpecIso_naturality, Iso.hom_inv_id_apply] + erw [Iso.hom_inv_id_apply] exact eval₂_X _ _ _ inv_hom_id := by apply ext_of_isAffine simp only [Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, - TopologicalSpace.Opens.map_top, Scheme.comp_app, Scheme.toSpecΓ_app_top, + TopologicalSpace.Opens.map_top, Scheme.comp_app, Scheme.toSpecΓ_appTop, Scheme.ΓSpecIso_naturality, Category.assoc, Scheme.id_app, ← Iso.eq_inv_comp, Category.comp_id] apply ringHom_ext' · show _ = CommRingCat.ofHom C ≫ _ rw [CommRingCat.comp_eq_ring_hom_comp, RingHom.comp_assoc, eval₂Hom_comp_C, ← CommRingCat.comp_eq_ring_hom_comp, ← cancel_mono (Scheme.ΓSpecIso _).hom] - erw [← Scheme.comp_app] - rw [homOfVector_over, Scheme.comp_app] - simp only [TopologicalSpace.Opens.map_top, Category.assoc, Scheme.ΓSpecIso_naturality, ← - Scheme.toSpecΓ_app_top] - erw [← Scheme.comp_app_assoc] - rw [Scheme.isoSpec, asIso_inv, IsIso.hom_inv_id] + rw [← Scheme.comp_appTop, homOfVector_over, Scheme.comp_appTop] + simp only [Category.assoc, Scheme.ΓSpecIso_naturality, ← Scheme.toSpecΓ_appTop] + rw [← Scheme.comp_appTop_assoc, Scheme.isoSpec, asIso_inv, IsIso.hom_inv_id] simp rfl · intro i erw [CommRingCat.comp_apply, coe_eval₂Hom] simp only [eval₂_X] - exact homOfVector_app_top_coord _ _ _ + exact homOfVector_appTop_coord _ _ _ @[simp] -lemma isoOfIsAffine_hom_app_top [IsAffine S] : - (isoOfIsAffine n S).hom.app ⊤ = - (Scheme.ΓSpecIso _).hom ≫ eval₂Hom ((𝔸(n; S) ↘ S).app ⊤) (coord S) := by +lemma isoOfIsAffine_hom_appTop [IsAffine S] : + (isoOfIsAffine n S).hom.appTop = + (Scheme.ΓSpecIso _).hom ≫ eval₂Hom ((𝔸(n; S) ↘ S).appTop) (coord S) := by simp [isoOfIsAffine_hom] -@[local simp] -lemma isoOfIsAffine_inv_app_top_coord [IsAffine S] (i) : - (isoOfIsAffine n S).inv.app ⊤ (coord _ i) = (Scheme.ΓSpecIso (.of _)).inv (.X i) := - homOfVector_app_top_coord _ _ _ +@[simp] +lemma isoOfIsAffine_inv_appTop_coord [IsAffine S] (i) : + (isoOfIsAffine n S).inv.appTop (coord _ i) = (Scheme.ΓSpecIso (.of _)).inv (.X i) := + homOfVector_appTop_coord _ _ _ @[reassoc (attr := simp)] lemma isoOfIsAffine_inv_over [IsAffine S] : @@ -238,27 +233,27 @@ def SpecIso (R : CommRingCat.{max u v}) : (Scheme.ΓSpecIso R).symm.commRingCatIsoToRingEquiv).toCommRingCatIso.op @[simp] -lemma SpecIso_hom_app_top (R : CommRingCat.{max u v}) : - (SpecIso n R).hom.app ⊤ = (Scheme.ΓSpecIso _).hom ≫ - eval₂Hom ((Scheme.ΓSpecIso _).inv ≫ (𝔸(n; Spec R) ↘ Spec R).app ⊤) (coord (Spec R)) := by +lemma SpecIso_hom_appTop (R : CommRingCat.{max u v}) : + (SpecIso n R).hom.appTop = (Scheme.ΓSpecIso _).hom ≫ + eval₂Hom ((Scheme.ΓSpecIso _).inv ≫ (𝔸(n; Spec R) ↘ Spec R).appTop) (coord (Spec R)) := by simp only [SpecIso, Iso.trans_hom, Functor.mapIso_hom, Iso.op_hom, RingEquiv.toCommRingCatIso_hom, RingEquiv.toRingHom_eq_coe, Scheme.Spec_map, Quiver.Hom.unop_op, Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, TopologicalSpace.Opens.map_top, Scheme.comp_app, - isoOfIsAffine_hom_app_top] + isoOfIsAffine_hom_appTop] erw [Scheme.ΓSpecIso_naturality_assoc] congr 1 apply ringHom_ext' - · ext; simp; rfl + · ext; simp · simp -@[local simp] -lemma SpecIso_inv_app_top_coord (R : CommRingCat.{max u v}) (i) : - (SpecIso n R).inv.app ⊤ (coord _ i) = (Scheme.ΓSpecIso (.of _)).inv (.X i) := by +@[simp] +lemma SpecIso_inv_appTop_coord (R : CommRingCat.{max u v}) (i) : + (SpecIso n R).inv.appTop (coord _ i) = (Scheme.ΓSpecIso (.of _)).inv (.X i) := by simp only [SpecIso, Iso.trans_inv, Functor.mapIso_inv, Iso.op_inv, RingEquiv.toCommRingCatIso_inv, mapEquiv_symm, RingEquiv.toRingHom_eq_coe, Scheme.Spec_map, Quiver.Hom.unop_op, Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, TopologicalSpace.Opens.map_top, Scheme.comp_app, CommRingCat.comp_apply] - erw [isoOfIsAffine_inv_app_top_coord, ← CommRingCat.comp_apply] + erw [isoOfIsAffine_inv_appTop_coord, ← CommRingCat.comp_apply] rw [← Scheme.ΓSpecIso_inv_naturality] erw [CommRingCat.comp_apply] congr 1 @@ -286,14 +281,14 @@ def map {S T : Scheme.{max u v}} (f : S ⟶ T) : 𝔸(n; S) ⟶ 𝔸(n; T) := lemma map_over {S T : Scheme.{max u v}} (f : S ⟶ T) : map n f ≫ 𝔸(n; T) ↘ T = 𝔸(n; S) ↘ S ≫ f := pullback.lift_fst _ _ _ -@[local simp] -lemma map_app_top_coord {S T : Scheme.{max u v}} (f : S ⟶ T) (i) : - (map n f).app ⊤ (coord T i) = coord S i := - homOfVector_app_top_coord _ _ _ +@[simp] +lemma map_appTop_coord {S T : Scheme.{max u v}} (f : S ⟶ T) (i) : + (map n f).appTop (coord T i) = coord S i := + homOfVector_appTop_coord _ _ _ @[simp] lemma map_id : map n (𝟙 S) = 𝟙 𝔸(n; S) := by - ext1 <;> simp [-TopologicalSpace.Opens.map_top]; rfl + ext1 <;> simp @[reassoc, simp] lemma map_comp {S S' S'' : Scheme} (f : S ⟶ S') (g : S' ⟶ S'') : @@ -302,7 +297,7 @@ lemma map_comp {S S' S'' : Scheme} (f : S ⟶ S') (g : S' ⟶ S'') : · simp · simp only [TopologicalSpace.Opens.map_top, Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, Scheme.comp_app, CommRingCat.comp_apply] - erw [map_app_top_coord, map_app_top_coord, map_app_top_coord] + erw [map_appTop_coord, map_appTop_coord, map_appTop_coord] lemma map_Spec_map {R S : CommRingCat.{max u v}} (φ : R ⟶ S) : map n (Spec.map φ) = @@ -314,9 +309,9 @@ lemma map_Spec_map {R S : CommRingCat.{max u v}} (φ : R ⟶ S) : rw [map_comp_C] · simp only [Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, TopologicalSpace.Opens.map_top, Scheme.comp_app, CommRingCat.comp_apply] - conv_lhs => enter[2]; tactic => exact map_app_top_coord _ _ - conv_rhs => enter[2]; tactic => exact SpecIso_inv_app_top_coord _ _ - erw [SpecIso_inv_app_top_coord, ← CommRingCat.comp_apply] + conv_lhs => enter[2]; tactic => exact map_appTop_coord _ _ + conv_rhs => enter[2]; tactic => exact SpecIso_inv_appTop_coord _ _ + erw [SpecIso_inv_appTop_coord, ← CommRingCat.comp_apply] rw [← Scheme.ΓSpecIso_inv_naturality] erw [CommRingCat.comp_apply, map_X] rfl @@ -337,39 +332,29 @@ lemma reindex_over {n m : Type v} (i : m → n) (S : Scheme.{max u v}) : reindex i S ≫ 𝔸(m; S) ↘ S = 𝔸(n; S) ↘ S := pullback.lift_fst _ _ _ -@[local simp] -lemma reindex_app_top_coord {n m : Type v} (i : m → n) (S : Scheme.{max u v}) (j : m) : - (reindex i S).app ⊤ (coord S j) = coord S (i j) := - homOfVector_app_top_coord _ _ _ +@[simp] +lemma reindex_appTop_coord {n m : Type v} (i : m → n) (S : Scheme.{max u v}) (j : m) : + (reindex i S).appTop (coord S j) = coord S (i j) := + homOfVector_appTop_coord _ _ _ @[simp] lemma reindex_id : reindex id S = 𝟙 𝔸(n; S) := by - ext1 <;> simp [-TopologicalSpace.Opens.map_top]; rfl + ext1 <;> simp +@[simp, reassoc] lemma reindex_comp {n₁ n₂ n₃ : Type v} (i : n₁ → n₂) (j : n₂ → n₃) (S : Scheme.{max u v}) : reindex (j ∘ i) S = reindex j S ≫ reindex i S := by have H₁ : reindex (j ∘ i) S ≫ 𝔸(n₁; S) ↘ S = (reindex j S ≫ reindex i S) ≫ 𝔸(n₁; S) ↘ S := by simp - have H₂ (k) : (reindex (j ∘ i) S).app ⊤ (coord S k) = - (reindex j S).app ⊤ ((reindex i S).app ⊤ (coord S k)) := by - rw [reindex_app_top_coord, reindex_app_top_coord, reindex_app_top_coord] + have H₂ (k) : (reindex (j ∘ i) S).appTop (coord S k) = + (reindex j S).appTop ((reindex i S).appTop (coord S k)) := by + rw [reindex_appTop_coord, reindex_appTop_coord, reindex_appTop_coord] rfl exact hom_ext H₁ H₂ --- These time out if added to `reindex_comp` directly. -attribute [reassoc] reindex_comp -attribute [simp] reindex_comp - +@[reassoc (attr := simp)] lemma map_reindex {n₁ n₂ : Type v} (i : n₁ → n₂) {S T : Scheme.{max u v}} (f : S ⟶ T) : map n₂ f ≫ reindex i T = reindex i S ≫ map n₁ f := by - apply hom_ext - · simp - · intro j - simp only [Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, Scheme.comp_app, - CommRingCat.comp_apply, map_app_top_coord, reindex_app_top_coord] - simp only [TopologicalSpace.Opens.map_top] - erw [map_app_top_coord f (i j), reindex_app_top_coord i S] - -attribute [reassoc (attr := simp)] map_reindex + apply hom_ext <;> simp /-- The affine space as a functor. -/ @[simps] diff --git a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean index 6c8faa057ea86..87195c4877f61 100644 --- a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean +++ b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean @@ -415,12 +415,12 @@ theorem Scheme.toSpecΓ_base (X : Scheme.{u}) (x) : @[reassoc (attr := simp)] theorem Scheme.toSpecΓ_naturality {X Y : Scheme.{u}} (f : X ⟶ Y) : - f ≫ Y.toSpecΓ = X.toSpecΓ ≫ Spec.map (f.app ⊤) := + f ≫ Y.toSpecΓ = X.toSpecΓ ≫ Spec.map (f.appTop) := ΓSpec.adjunction.unit.naturality f @[simp] -theorem Scheme.toSpecΓ_app_top (X : Scheme.{u}) : - X.toSpecΓ.app ⊤ = (Scheme.ΓSpecIso Γ(X, ⊤)).hom := by +theorem Scheme.toSpecΓ_appTop (X : Scheme.{u}) : + X.toSpecΓ.appTop = (Scheme.ΓSpecIso Γ(X, ⊤)).hom := by have := ΓSpec.adjunction.left_triangle_components X dsimp at this rw [← IsIso.eq_comp_inv] at this @@ -428,6 +428,8 @@ theorem Scheme.toSpecΓ_app_top (X : Scheme.{u}) : Scheme.Γ_obj, Category.id_comp] at this rw [← Quiver.Hom.op_inj.eq_iff, this, ← op_inv, IsIso.Iso.inv_inv] +@[deprecated (since := "2024-11-23")] alias Scheme.toSpecΓ_app_top := Scheme.toSpecΓ_appTop + @[simp] theorem SpecMap_ΓSpecIso_hom (R : CommRingCat.{u}) : Spec.map ((Scheme.ΓSpecIso R).hom) = (Spec R).toSpecΓ := by @@ -437,9 +439,9 @@ theorem SpecMap_ΓSpecIso_hom (R : CommRingCat.{u}) : 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] + rw [← basicOpen_eq_of_affine, Scheme.preimage_basicOpen, ← Scheme.Hom.appTop] congr - rw [Scheme.toSpecΓ_app_top] + rw [Scheme.toSpecΓ_appTop] exact Iso.inv_hom_id_apply _ _ -- Warning: this LHS of this lemma breaks the structure-sheaf abstraction. @@ -456,25 +458,24 @@ theorem toOpen_toSpecΓ_app {X : Scheme.{u}} (U) : exact Category.id_comp _ lemma ΓSpecIso_inv_ΓSpec_adjunction_homEquiv {X : Scheme.{u}} {B : CommRingCat} (φ : B ⟶ Γ(X, ⊤)) : - (Scheme.ΓSpecIso B).inv ≫ ((ΓSpec.adjunction.homEquiv X (op B)) φ.op).app ⊤ = φ := by + (Scheme.ΓSpecIso B).inv ≫ ((ΓSpec.adjunction.homEquiv X (op B)) φ.op).appTop = φ := by simp only [Adjunction.homEquiv_apply, Scheme.Spec_map, Opens.map_top, Scheme.comp_app] simp lemma ΓSpec_adjunction_homEquiv_eq {X : Scheme.{u}} {B : CommRingCat} (φ : B ⟶ Γ(X, ⊤)) : - (((ΓSpec.adjunction.homEquiv X (op B)) φ.op).app ⊤) = (Scheme.ΓSpecIso B).hom ≫ φ := by - simp_rw [← ΓSpecIso_inv_ΓSpec_adjunction_homEquiv φ] - simp + ((ΓSpec.adjunction.homEquiv X (op B)) φ.op).appTop = (Scheme.ΓSpecIso B).hom ≫ φ := by + rw [← Iso.inv_comp_eq, ΓSpecIso_inv_ΓSpec_adjunction_homEquiv] theorem ΓSpecIso_obj_hom {X : Scheme.{u}} (U : X.Opens) : - (Scheme.ΓSpecIso Γ(X, U)).hom = (Spec.map U.topIso.inv).app ⊤ ≫ - U.toScheme.toSpecΓ.app ⊤ ≫ U.topIso.hom := by simp + (Scheme.ΓSpecIso Γ(X, U)).hom = (Spec.map U.topIso.inv).appTop ≫ + U.toScheme.toSpecΓ.appTop ≫ 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 +alias ΓSpec.adjunction_unit_app_app_top := Scheme.toSpecΓ_appTop @[deprecated (since := "2024-07-24")] alias ΓSpec.adjunction_unit_map_basicOpen := Scheme.toSpecΓ_preimage_basicOpen diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean b/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean index 66b28ced66e1b..10aa7de982be9 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean @@ -158,7 +158,7 @@ instance : HasAffineProperty @IsAffineHom fun X _ _ _ ↦ IsAffine X where rw [Scheme.preimage_basicOpen] exact (isAffineOpen_top X).basicOpen _ · intro X Y _ f S hS hS' - apply_fun Ideal.map (f.app ⊤) at hS + apply_fun Ideal.map (f.appTop) at hS rw [Ideal.map_span, Ideal.map_top] at hS apply isAffine_of_isAffineOpen_basicOpen _ hS have : ∀ i : S, IsAffineOpen (f⁻¹ᵁ Y.basicOpen i.1) := hS' diff --git a/Mathlib/AlgebraicGeometry/Morphisms/AffineAnd.lean b/Mathlib/AlgebraicGeometry/Morphisms/AffineAnd.lean index 6eea058bd4042..632a366ca7236 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/AffineAnd.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/AffineAnd.lean @@ -35,11 +35,11 @@ variable (Q : ∀ {R S : Type u} [CommRing R] [CommRing S], (R →+* S) → Prop /-- This is the affine target morphism property where the source is affine and the induced map of rings on global sections satisfies `P`. -/ def affineAnd : AffineTargetMorphismProperty := - fun X _ f ↦ IsAffine X ∧ Q (f.app ⊤) + fun X _ f ↦ IsAffine X ∧ Q (f.appTop) @[simp] lemma affineAnd_apply {X Y : Scheme.{u}} (f : X ⟶ Y) [IsAffine Y] : - affineAnd Q f ↔ IsAffine X ∧ Q (f.app ⊤) := + affineAnd Q f ↔ IsAffine X ∧ Q (f.appTop) := Iff.rfl attribute [local simp] AffineTargetMorphismProperty.toProperty_apply @@ -64,28 +64,29 @@ lemma affineAnd_isLocal (hPi : RingHom.RespectsIso Q) (hQl : RingHom.Localizatio constructor · simp only [Scheme.preimage_basicOpen, Opens.map_top] exact (isAffineOpen_top X).basicOpen _ - · dsimp only [Opens.map_top] - rw [morphismRestrict_app, hPi.cancel_right_isIso, Scheme.Opens.ι_image_top] + · dsimp only + rw [morphismRestrict_appTop, hPi.cancel_right_isIso, Scheme.Opens.ι_image_top] rw [(isAffineOpen_top Y).app_basicOpen_eq_away_map f (isAffineOpen_top X), - hPi.cancel_right_isIso] - haveI := (isAffineOpen_top X).isLocalization_basicOpen (f.app ⊤ r) + hPi.cancel_right_isIso, ← Scheme.Hom.appTop] + dsimp only [Opens.map_top] + haveI := (isAffineOpen_top X).isLocalization_basicOpen (f.appTop r) apply hQl exact hf of_basicOpenCover {X Y _} f s hs hf := by dsimp [affineAnd] at hf haveI : IsAffine X := by - apply isAffine_of_isAffineOpen_basicOpen (f.app ⊤ '' s) - · apply_fun Ideal.map (f.app ⊤) at hs + apply isAffine_of_isAffineOpen_basicOpen (f.appTop '' s) + · apply_fun Ideal.map (f.appTop) at hs rwa [Ideal.map_span, Ideal.map_top] at hs · rintro - ⟨r, hr, rfl⟩ simp_rw [Scheme.preimage_basicOpen] at hf exact (hf ⟨r, hr⟩).left - refine ⟨inferInstance, hQs.ofIsLocalization' hPi (f.app ⊤) s hs fun a ↦ ?_⟩ - refine ⟨Γ(Y, Y.basicOpen a.val), Γ(X, X.basicOpen (f.app ⊤ a.val)), inferInstance, + refine ⟨inferInstance, hQs.ofIsLocalization' hPi (f.appTop) s hs fun a ↦ ?_⟩ + refine ⟨Γ(Y, Y.basicOpen a.val), Γ(X, X.basicOpen (f.appTop a.val)), inferInstance, inferInstance, inferInstance, inferInstance, inferInstance, ?_, ?_⟩ - · exact (isAffineOpen_top X).isLocalization_basicOpen (f.app ⊤ a.val) + · exact (isAffineOpen_top X).isLocalization_basicOpen (f.appTop a.val) · obtain ⟨_, hf⟩ := hf a - rw [morphismRestrict_app, hPi.cancel_right_isIso, Scheme.Opens.ι_image_top] at hf + rw [morphismRestrict_appTop, hPi.cancel_right_isIso, Scheme.Opens.ι_image_top] at hf rw [(isAffineOpen_top Y).app_basicOpen_eq_away_map _ (isAffineOpen_top X)] at hf rwa [hPi.cancel_right_isIso] at hf @@ -96,7 +97,7 @@ lemma affineAnd_isStableUnderBaseChange (hQi : RingHom.RespectsIso Q) haveI : (affineAnd Q).toProperty.RespectsIso := affineAnd_respectsIso hQi apply AffineTargetMorphismProperty.IsStableUnderBaseChange.mk intro X Y S _ _ f g ⟨hY, hg⟩ - exact ⟨inferInstance, hQb.pullback_fst_app_top _ hQi f _ hg⟩ + exact ⟨inferInstance, hQb.pullback_fst_appTop _ hQi f _ hg⟩ lemma targetAffineLocally_affineAnd_iff (hQi : RingHom.RespectsIso Q) {X Y : Scheme.{u}} (f : X ⟶ Y) : @@ -133,7 +134,7 @@ lemma targetAffineLocally_affineAnd_iff_affineLocally (hQ : RingHom.PropertyIsLo intro U have : IsAffine (f ⁻¹ᵁ U) := hf.isAffine_preimage U U.2 rw [HasRingHomProperty.iff_of_isAffine (P := affineLocally Q), - morphismRestrict_app, hQ.respectsIso.cancel_right_isIso] + morphismRestrict_appTop, hQ.respectsIso.cancel_right_isIso] apply h rw [Scheme.Opens.ι_image_top] exact U.2 diff --git a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean index d72195178231c..0145a3045d5db 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean @@ -117,7 +117,7 @@ instance spec_of_quotient_mk {R : CommRingCat.{u}} (I : Ideal R) : /-- Any morphism between affine schemes that is surjective on global sections is a closed immersion. -/ lemma of_surjective_of_isAffine {X Y : Scheme} [IsAffine X] [IsAffine Y] (f : X ⟶ Y) - (h : Function.Surjective (f.app ⊤)) : IsClosedImmersion f := by + (h : Function.Surjective (f.appTop)) : IsClosedImmersion f := by rw [MorphismProperty.arrow_mk_iso_iff @IsClosedImmersion (arrowIsoSpecΓOfIsAffine f)] apply spec_of_surjective exact h @@ -161,7 +161,7 @@ open IsClosedImmersion LocallyRingedSpace has a closed image and `f` induces an injection on global sections, then `f` is surjective. -/ lemma surjective_of_isClosed_range_of_injective [CompactSpace X] - (hfcl : IsClosed (Set.range f.base)) (hfinj : Function.Injective (f.app ⊤)) : + (hfcl : IsClosed (Set.range f.base)) (hfinj : Function.Injective (f.appTop)) : Function.Surjective f.base := by obtain ⟨I, hI⟩ := (Scheme.eq_zeroLocus_of_isClosed_of_isAffine Y (Set.range f.base)).mp hfcl let 𝒰 : X.OpenCover := X.affineCover.finiteSubcover @@ -186,12 +186,12 @@ lemma surjective_of_isClosed_range_of_injective [CompactSpace X] injective if it is injective on global sections. -/ lemma stalkMap_injective_of_isOpenMap_of_injective [CompactSpace X] (hfopen : IsOpenMap f.base) (hfinj₁ : Function.Injective f.base) - (hfinj₂ : Function.Injective (f.app ⊤)) (x : X) : + (hfinj₂ : Function.Injective (f.appTop)) (x : X) : Function.Injective (f.stalkMap x) := by - let φ : Γ(Y, ⊤) ⟶ Γ(X, ⊤) := f.app ⊤ + let φ : Γ(Y, ⊤) ⟶ Γ(X, ⊤) := f.appTop let 𝒰 : X.OpenCover := X.affineCover.finiteSubcover have (i : 𝒰.J) : IsAffine (𝒰.obj i) := Scheme.isAffine_affineCover X _ - let res (i : 𝒰.J) : Γ(X, ⊤) ⟶ Γ(𝒰.obj i, ⊤) := (𝒰.map i).app ⊤ + let res (i : 𝒰.J) : Γ(X, ⊤) ⟶ Γ(𝒰.obj i, ⊤) := (𝒰.map i).appTop refine stalkMap_injective_of_isAffine _ _ (fun (g : Γ(Y, ⊤)) h ↦ ?_) rw [TopCat.Presheaf.Γgerm, Scheme.stalkMap_germ_apply] at h obtain ⟨U, w, (hx : x ∈ U), hg⟩ := @@ -200,8 +200,8 @@ lemma stalkMap_injective_of_isOpenMap_of_injective [CompactSpace X] (show f.base x ∈ ⟨f.base '' U.carrier, hfopen U.carrier U.is_open'⟩ from ⟨x, by simpa⟩) let W (i : 𝒰.J) : TopologicalSpace.Opens (𝒰.obj i) := (𝒰.obj i).basicOpen ((res i) (φ s)) have hwle (i : 𝒰.J) : W i ≤ (𝒰.map i)⁻¹ᵁ U := by - show ((𝒰.obj i).basicOpen ((𝒰.map i ≫ f).app ⊤ s)) ≤ _ - rw [← Scheme.preimage_basicOpen, Scheme.comp_coeBase, Opens.map_comp_obj] + show (𝒰.obj i).basicOpen ((𝒰.map i ≫ f).appTop s) ≤ _ + rw [← Scheme.preimage_basicOpen_top, Scheme.comp_coeBase, Opens.map_comp_obj] refine Scheme.Hom.preimage_le_preimage_of_le _ (le_trans (f.preimage_le_preimage_of_le bsle) (le_of_eq ?_)) simp [Set.preimage_image_eq _ hfinj₁] @@ -229,7 +229,7 @@ namespace IsClosedImmersion /-- If `f` is a closed immersion with affine target such that the induced map on global sections is injective, `f` is an isomorphism. -/ theorem isIso_of_injective_of_isAffine [IsClosedImmersion f] - (hf : Function.Injective (f.app ⊤)) : IsIso f := (isIso_iff_stalk_iso f).mpr <| + (hf : Function.Injective (f.appTop)) : IsIso f := (isIso_iff_stalk_iso f).mpr <| have : CompactSpace X := f.isClosedEmbedding.compactSpace have hiso : IsIso f.base := TopCat.isIso_of_bijective_of_isClosedMap _ ⟨f.isClosedEmbedding.injective, @@ -244,7 +244,7 @@ variable (f) /-- If `f` is a closed immersion with affine target, the source is affine and the induced map on global sections is surjective. -/ theorem isAffine_surjective_of_isAffine [IsClosedImmersion f] : - IsAffine X ∧ Function.Surjective (f.app ⊤) := by + IsAffine X ∧ Function.Surjective (f.appTop) := by haveI i : IsClosedImmersion f := inferInstance rw [← affineTargetImageFactorization_comp f] at i ⊢ haveI := of_surjective_of_isAffine (affineTargetImageInclusion f) @@ -254,7 +254,7 @@ theorem isAffine_surjective_of_isAffine [IsClosedImmersion f] : haveI := isIso_of_injective_of_isAffine (affineTargetImageFactorization_app_injective f) exact ⟨isAffine_of_isIso (affineTargetImageFactorization f), (ConcreteCategory.bijective_of_isIso - ((affineTargetImageFactorization f).app ⊤)).surjective.comp <| + ((affineTargetImageFactorization f).appTop)).surjective.comp <| affineTargetImageInclusion_app_surjective f⟩ end IsClosedImmersion @@ -268,7 +268,7 @@ instance IsClosedImmersion.isLocalAtTarget : IsLocalAtTarget @IsClosedImmersion /-- On morphisms with affine target, being a closed immersion is precisely having affine source and being surjective on global sections. -/ instance IsClosedImmersion.hasAffineProperty : HasAffineProperty @IsClosedImmersion - (fun X _ f ↦ IsAffine X ∧ Function.Surjective (f.app ⊤)) := by + (fun X _ f ↦ IsAffine X ∧ Function.Surjective (f.appTop)) := by convert HasAffineProperty.of_isLocalAtTarget @IsClosedImmersion refine ⟨fun ⟨h₁, h₂⟩ ↦ of_surjective_of_isAffine _ h₂, by apply isAffine_surjective_of_isAffine⟩ @@ -290,7 +290,7 @@ instance IsClosedImmersion.isStableUnderBaseChange : haveI := HasAffineProperty.isLocal_affineProperty @IsClosedImmersion apply AffineTargetMorphismProperty.IsStableUnderBaseChange.mk intro X Y S _ _ f g ⟨ha, hsurj⟩ - exact ⟨inferInstance, RingHom.surjective_isStableUnderBaseChange.pullback_fst_app_top _ + exact ⟨inferInstance, RingHom.surjective_isStableUnderBaseChange.pullback_fst_appTop _ RingHom.surjective_respectsIso f _ hsurj⟩ /-- Closed immersions are locally of finite type. -/ @@ -320,13 +320,13 @@ lemma isIso_of_isClosedImmersion_of_surjective {X Y : Scheme.{u}} (f : X ⟶ Y) · infer_instance apply IsClosedImmersion.isIso_of_injective_of_isAffine obtain ⟨hX, hf⟩ := HasAffineProperty.iff_of_isAffine.mp ‹IsClosedImmersion f› - let φ := f.app ⊤ + let φ := f.appTop suffices RingHom.ker φ ≤ nilradical _ by rwa [nilradical_eq_zero, Submodule.zero_eq_bot, le_bot_iff, ← RingHom.injective_iff_ker_eq_bot] at this refine (PrimeSpectrum.zeroLocus_eq_top_iff _).mp ?_ rw [← range_specComap_of_surjective _ _ hf, Set.top_eq_univ, Set.range_eq_univ] - have : Surjective (Spec.map (f.app ⊤)) := + have : Surjective (Spec.map (f.appTop)) := (MorphismProperty.arrow_mk_iso_iff @Surjective (arrowIsoSpecΓOfIsAffine f)).mp (inferInstanceAs (Surjective f)) exact this.1 diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Finite.lean b/Mathlib/AlgebraicGeometry/Morphisms/Finite.lean index 0f054e90a059e..60115fff857f2 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Finite.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Finite.lean @@ -37,7 +37,7 @@ class IsFinite {X Y : Scheme} (f : X ⟶ Y) extends IsAffineHom f : Prop where namespace IsFinite -instance : HasAffineProperty @IsFinite (fun X _ f _ ↦ IsAffine X ∧ RingHom.Finite (f.app ⊤)) := by +instance : HasAffineProperty @IsFinite (fun X _ f _ ↦ IsAffine X ∧ RingHom.Finite (f.appTop)) := by show HasAffineProperty @IsFinite (affineAnd RingHom.Finite) rw [HasAffineProperty.affineAnd_iff _ RingHom.finite_respectsIso RingHom.finite_localizationPreserves RingHom.finite_ofLocalizationSpan] diff --git a/Mathlib/AlgebraicGeometry/Morphisms/IsIso.lean b/Mathlib/AlgebraicGeometry/Morphisms/IsIso.lean index 06cb129fa9785..ed9fb2ee4ca4d 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/IsIso.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/IsIso.lean @@ -36,10 +36,10 @@ lemma isomorphisms_eq_stalkwise : instance : IsLocalAtTarget (isomorphisms Scheme) := isomorphisms_eq_isOpenImmersion_inf_surjective ▸ inferInstance -instance : HasAffineProperty (isomorphisms Scheme) fun X _ f _ ↦ IsAffine X ∧ IsIso (f.app ⊤) := by +instance : HasAffineProperty (isomorphisms Scheme) fun X _ f _ ↦ IsAffine X ∧ IsIso (f.appTop) := by convert HasAffineProperty.of_isLocalAtTarget (isomorphisms Scheme) with X Y f hY exact ⟨fun ⟨_, _⟩ ↦ (arrow_mk_iso_iff (isomorphisms _) (arrowIsoSpecΓOfIsAffine f)).mpr - (inferInstanceAs (IsIso (Spec.map (f.app ⊤)))), + (inferInstanceAs (IsIso (Spec.map (f.appTop)))), fun (_ : IsIso f) ↦ ⟨isAffine_of_isIso f, inferInstance⟩⟩ instance : IsLocalAtTarget (monomorphisms Scheme) := diff --git a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean index 6a02465a5cb1a..4306631e05eb2 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean @@ -39,9 +39,9 @@ For `HasRingHomProperty P Q` and `f : X ⟶ Y`, we provide these API lemmas: - `AlgebraicGeometry.HasRingHomProperty.iff_appLE`: `P f` if and only if `Q (f.appLE U V _)` for all affine `U : Opens Y` and `V : Opens X`. - `AlgebraicGeometry.HasRingHomProperty.iff_of_source_openCover`: - If `Y` is affine, `P f ↔ ∀ i, Q ((𝒰.map i ≫ f).app ⊤)` for an affine open cover `𝒰` of `X`. + If `Y` is affine, `P f ↔ ∀ i, Q ((𝒰.map i ≫ f).appTop)` for an affine open cover `𝒰` of `X`. - `AlgebraicGeometry.HasRingHomProperty.iff_of_isAffine`: - If `X` and `Y` are affine, then `P f ↔ Q (f.app ⊤)`. + If `X` and `Y` are affine, then `P f ↔ Q (f.appTop)`. - `AlgebraicGeometry.HasRingHomProperty.Spec_iff`: `P (Spec.map φ) ↔ Q φ` - `AlgebraicGeometry.HasRingHomProperty.iff_of_iSup_eq_top`: @@ -66,14 +66,14 @@ namespace RingHom variable (P : ∀ {R S : Type u} [CommRing R] [CommRing S], (R →+* S) → Prop) -theorem IsStableUnderBaseChange.pullback_fst_app_top +theorem IsStableUnderBaseChange.pullback_fst_appTop (hP : IsStableUnderBaseChange P) (hP' : RespectsIso P) {X Y S : Scheme} [IsAffine X] [IsAffine Y] [IsAffine S] (f : X ⟶ S) (g : Y ⟶ S) - (H : P (g.app ⊤)) : P ((pullback.fst f g).app ⊤) := by + (H : P g.appTop) : P (pullback.fst f g).appTop := by -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11224): change `rw` to `erw` erw [← PreservesPullback.iso_inv_fst AffineScheme.forgetToScheme (AffineScheme.ofHom f) (AffineScheme.ofHom g)] - rw [Scheme.comp_app, hP'.cancel_right_isIso, AffineScheme.forgetToScheme_map] + rw [Scheme.comp_appTop, hP'.cancel_right_isIso, AffineScheme.forgetToScheme_map] have := congr_arg Quiver.Hom.unop (PreservesPullback.iso_hom_fst AffineScheme.Γ.rightOp (AffineScheme.ofHom f) (AffineScheme.ofHom g)) @@ -83,6 +83,10 @@ theorem IsStableUnderBaseChange.pullback_fst_app_top rw [← this, hP'.cancel_right_isIso, ← pushoutIsoUnopPullback_inl_hom, hP'.cancel_right_isIso] exact hP.pushout_inl _ hP' _ _ H +@[deprecated (since := "2024-11-23")] +alias IsStableUnderBaseChange.pullback_fst_app_top := +IsStableUnderBaseChange.pullback_fst_appTop + end RingHom namespace AlgebraicGeometry @@ -268,10 +272,12 @@ theorem appLE (H : P f) (U : Y.affineOpens) (V : X.affineOpens) (e) : Q (f.appLE rw [eq_affineLocally P, affineLocally_iff_affineOpens_le] at H exact H _ _ _ -theorem app_top (H : P f) [IsAffine X] [IsAffine Y] : Q (f.app ⊤) := by - rw [Scheme.Hom.app_eq_appLE] +theorem appTop (H : P f) [IsAffine X] [IsAffine Y] : Q f.appTop := by + rw [Scheme.Hom.appTop, Scheme.Hom.app_eq_appLE] exact appLE P f H ⟨_, isAffineOpen_top _⟩ ⟨_, isAffineOpen_top _⟩ _ +@[deprecated (since := "2024-11-23")] alias app_top := appTop + include Q in theorem comp_of_isOpenImmersion [IsOpenImmersion f] (H : P g) : P (f ≫ g) := by @@ -291,7 +297,7 @@ lemma iff_appLE : P f ↔ ∀ (U : Y.affineOpens) (V : X.affineOpens) (e), Q (f. rw [eq_affineLocally P, affineLocally_iff_affineOpens_le] theorem of_source_openCover [IsAffine Y] - (𝒰 : X.OpenCover) [∀ i, IsAffine (𝒰.obj i)] (H : ∀ i, Q ((𝒰.map i ≫ f).app ⊤)) : + (𝒰 : X.OpenCover) [∀ i, IsAffine (𝒰.obj i)] (H : ∀ i, Q ((𝒰.map i ≫ f).appTop)) : P f := by rw [HasAffineProperty.iff_of_isAffine (P := P)] intro U @@ -313,17 +319,17 @@ theorem of_source_openCover [IsAffine Y] specialize H i rw [← (isLocal_ringHomProperty P).respectsIso.cancel_right_isIso _ ((IsOpenImmersion.isoOfRangeEq (𝒰.map i) (S i).1.ι - Subtype.range_coe.symm).inv.app _), ← Scheme.comp_app, - IsOpenImmersion.isoOfRangeEq_inv_fac_assoc, Scheme.comp_app, - Scheme.Opens.ι_app, Scheme.Hom.app_eq_appLE, Scheme.Hom.appLE_map] at H + Subtype.range_coe.symm).inv.app _), ← Scheme.comp_appTop, + IsOpenImmersion.isoOfRangeEq_inv_fac_assoc, Scheme.comp_appTop, + Scheme.Opens.ι_appTop, Scheme.Hom.appTop, Scheme.Hom.app_eq_appLE, Scheme.Hom.appLE_map] at H exact (f.appLE_congr _ rfl (by simp) Q).mp H theorem iff_of_source_openCover [IsAffine Y] (𝒰 : X.OpenCover) [∀ i, IsAffine (𝒰.obj i)] : - P f ↔ ∀ i, Q ((𝒰.map i ≫ f).app ⊤) := - ⟨fun H i ↦ app_top P _ (comp_of_isOpenImmersion P (𝒰.map i) f H), of_source_openCover 𝒰⟩ + P f ↔ ∀ i, Q ((𝒰.map i ≫ f).appTop) := + ⟨fun H i ↦ appTop P _ (comp_of_isOpenImmersion P (𝒰.map i) f H), of_source_openCover 𝒰⟩ theorem iff_of_isAffine [IsAffine X] [IsAffine Y] : - P f ↔ Q (f.app ⊤) := by + P f ↔ Q (f.appTop) := by rw [iff_of_source_openCover (P := P) (Scheme.coverOfIsIso.{u} (𝟙 _))] simp @@ -359,7 +365,7 @@ lemma containsIdentities (hP : RingHom.ContainsIdentities Q) : P.ContainsIdentit rw [IsLocalAtTarget.iff_of_iSup_eq_top (P := P) _ (iSup_affineOpens_eq_top _)] intro U have : IsAffine (𝟙 X ⁻¹ᵁ U.1) := U.2 - rw [morphismRestrict_id, iff_of_isAffine (P := P), Scheme.id_app] + rw [morphismRestrict_id, iff_of_isAffine (P := P), Scheme.id_appTop] apply hP variable (P) in @@ -504,7 +510,7 @@ lemma isStableUnderBaseChange (hP : RingHom.IsStableUnderBaseChange Q) : limit.lift_π, PullbackCone.mk_pt, PullbackCone.mk_π_app, Category.comp_id] apply this _ (comp_of_isOpenImmersion _ _ _ H) inferInstance rw [iff_of_isAffine (P := P)] at H ⊢ - exact hP.pullback_fst_app_top _ (isLocal_ringHomProperty P).respectsIso _ _ H + exact hP.pullback_fst_appTop _ (isLocal_ringHomProperty P).respectsIso _ _ H include Q in private lemma respects_isOpenImmersion_aux diff --git a/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean b/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean index 548f463af6395..66aff1bb5a74d 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean @@ -74,7 +74,6 @@ variable {f} in lemma iff_of_isAffine [IsAffine X] [IsAffine Y] : SurjectiveOnStalks f ↔ RingHom.SurjectiveOnStalks (f.app ⊤) := by rw [← Spec_iff, MorphismProperty.arrow_mk_iso_iff @SurjectiveOnStalks (arrowIsoSpecΓOfIsAffine f)] - rfl theorem of_comp [SurjectiveOnStalks (f ≫ g)] : SurjectiveOnStalks f := by refine ⟨fun x ↦ ?_⟩ diff --git a/Mathlib/AlgebraicGeometry/Restrict.lean b/Mathlib/AlgebraicGeometry/Restrict.lean index 4977b2dfaeffc..1d336822caba3 100644 --- a/Mathlib/AlgebraicGeometry/Restrict.lean +++ b/Mathlib/AlgebraicGeometry/Restrict.lean @@ -70,6 +70,11 @@ lemma ι_app (V) : U.ι.app V = X.presheaf.map (homOfLE (x := U.ι ''ᵁ U.ι ⁻¹ᵁ V) (Set.image_preimage_subset _ _)).op := rfl +@[simp] +lemma ι_appTop : + U.ι.appTop = X.presheaf.map (homOfLE (x := U.ι ''ᵁ ⊤) le_top).op := + rfl + @[simp] lemma ι_appLE (V W e) : U.ι.appLE V W e = @@ -170,7 +175,7 @@ def opensRestrict : instance ΓRestrictAlgebra {X : Scheme.{u}} (U : X.Opens) : Algebra (Γ(X, ⊤)) Γ(U, ⊤) := - (U.ι.app ⊤).toAlgebra + U.ι.appTop.toAlgebra lemma Scheme.map_basicOpen (r : Γ(U, ⊤)) : U.ι ''ᵁ U.toScheme.basicOpen r = X.basicOpen @@ -246,6 +251,11 @@ theorem Scheme.homOfLE_app {U V : X.Opens} (e : U ≤ V) (W : Opens V) : rw [e₃, ← Functor.map_comp] congr 1 +theorem Scheme.homOfLE_appTop {U V : X.Opens} (e : U ≤ V) : + (X.homOfLE e).appTop = + X.presheaf.map (homOfLE <| X.ι_image_homOfLE_le_ι_image e ⊤).op := + homOfLE_app .. + instance (X : Scheme.{u}) {U V : X.Opens} (e : U ≤ V) : IsOpenImmersion (X.homOfLE e) := by delta Scheme.homOfLE infer_instance @@ -295,7 +305,7 @@ def Scheme.restrictFunctorΓ : X.restrictFunctor.op ⋙ (Over.forget X).op ⋙ S (by intro U V i dsimp - rw [X.homOfLE_app, ← Functor.map_comp, ← Functor.map_comp] + rw [X.homOfLE_appTop, ← Functor.map_comp, ← Functor.map_comp] congr 1) /-- `X ∣_ U ∣_ V` is isomorphic to `X ∣_ V ∣_ U` -/ @@ -531,6 +541,11 @@ theorem morphismRestrict_app {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (V : Quiver.Hom.unop_op, Hom.opensFunctor_map_homOfLE] rw [this, Hom.appLE_map, Hom.appLE_map, Hom.appLE_map] +theorem morphismRestrict_appTop {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) : + (f ∣_ U).appTop = f.app (U.ι ''ᵁ ⊤) ≫ + X.presheaf.map (eqToHom (image_morphismRestrict_preimage f U ⊤)).op := + morphismRestrict_app .. + @[simp] theorem morphismRestrict_app' {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (V : Opens U) : (f ∣_ U).app V = f.appLE _ _ (image_morphismRestrict_preimage f U V).le := @@ -547,7 +562,7 @@ theorem Γ_map_morphismRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) : Scheme.Γ.map (f ∣_ U).op = Y.presheaf.map (eqToHom U.isOpenEmbedding_obj_top.symm).op ≫ f.app U ≫ X.presheaf.map (eqToHom (f ⁻¹ᵁ U).isOpenEmbedding_obj_top).op := by - rw [Scheme.Γ_map_op, morphismRestrict_app f U ⊤, f.naturality_assoc, ← X.presheaf.map_comp] + rw [Scheme.Γ_map_op, morphismRestrict_appTop f U, f.naturality_assoc, ← X.presheaf.map_comp] rfl /-- Restricting a morphism onto the image of an open immersion is isomorphic to the base change @@ -694,11 +709,11 @@ end Scheme.Hom /-- `f.resLE U V` induces `f.appLE U V` on global sections. -/ noncomputable def arrowResLEAppIso (f : X ⟶ Y) (U : Y.Opens) (V : X.Opens) (e : V ≤ f ⁻¹ᵁ U) : - Arrow.mk ((f.resLE U V e).app ⊤) ≅ Arrow.mk (f.appLE U V e) := + Arrow.mk ((f.resLE U V e).appTop) ≅ Arrow.mk (f.appLE U V e) := Arrow.isoMk U.topIso V.topIso <| by simp only [Opens.map_top, Arrow.mk_left, Arrow.mk_right, Functor.id_obj, Scheme.Opens.topIso_hom, eqToHom_op, Arrow.mk_hom, Scheme.Hom.map_appLE] - rw [← Scheme.Hom.appLE_eq_app, Scheme.Hom.resLE_appLE, Scheme.Hom.appLE_map] + rw [Scheme.Hom.appTop, ← Scheme.Hom.appLE_eq_app, Scheme.Hom.resLE_appLE, Scheme.Hom.appLE_map] end MorphismRestrict diff --git a/Mathlib/AlgebraicGeometry/Scheme.lean b/Mathlib/AlgebraicGeometry/Scheme.lean index 6decb4e3cb660..dc282d80303ad 100644 --- a/Mathlib/AlgebraicGeometry/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/Scheme.lean @@ -108,6 +108,11 @@ this is the induced map `Γ(Y, U) ⟶ Γ(X, f ⁻¹ᵁ U)`. -/ abbrev app (U : Y.Opens) : Γ(Y, U) ⟶ Γ(X, f ⁻¹ᵁ U) := f.c.app (op U) +/-- Given a morphism of schemes `f : X ⟶ Y`, +this is the induced map `Γ(Y, ⊤) ⟶ Γ(X, ⊤)`. -/ +abbrev appTop : Γ(Y, ⊤) ⟶ Γ(X, ⊤) := + f.app ⊤ + @[reassoc] lemma naturality (i : op U' ⟶ op U) : Y.presheaf.map i ≫ f.app U = f.app U' ≫ X.presheaf.map ((Opens.map f.base).map i.unop).op := @@ -250,6 +255,11 @@ theorem id.base (X : Scheme) : (𝟙 X : _).base = 𝟙 _ := theorem id_app {X : Scheme} (U : X.Opens) : (𝟙 X : _).app U = 𝟙 _ := rfl +@[simp] +theorem id_appTop {X : Scheme} : + (𝟙 X : _).appTop = 𝟙 _ := + rfl + @[reassoc] theorem comp_toLRSHom {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) : (f ≫ g).toLRSHom = f.toLRSHom ≫ g.toLRSHom := @@ -275,6 +285,11 @@ theorem comp_app {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) (U) : (f ≫ g).app U = g.app U ≫ f.app _ := rfl +@[simp, reassoc] -- reassoc lemma does not need `simp` +theorem comp_appTop {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) : + (f ≫ g).appTop = g.appTop ≫ f.appTop := + rfl + @[deprecated (since := "2024-06-23")] alias comp_val_c_app := comp_app @[deprecated (since := "2024-06-23")] alias comp_val_c_app_assoc := comp_app_assoc @@ -335,8 +350,10 @@ theorem inv_app {X Y : Scheme} (f : X ⟶ Y) [IsIso f] (U : X.Opens) : rw [IsIso.eq_comp_inv, ← Scheme.comp_app, Scheme.congr_app (IsIso.hom_inv_id f), Scheme.id_app, Category.id_comp] -theorem inv_app_top {X Y : Scheme} (f : X ⟶ Y) [IsIso f] : - (inv f).app ⊤ = inv (f.app ⊤) := by simp +theorem inv_appTop {X Y : Scheme} (f : X ⟶ Y) [IsIso f] : + (inv f).appTop = inv (f.appTop) := by simp + +@[deprecated (since := "2024-11-23")] alias inv_app_top := inv_appTop end Scheme @@ -440,10 +457,10 @@ theorem Γ_obj_op (X : Scheme) : Γ.obj (op X) = Γ(X, ⊤) := rfl @[simp] -theorem Γ_map {X Y : Schemeᵒᵖ} (f : X ⟶ Y) : Γ.map f = f.unop.app ⊤ := +theorem Γ_map {X Y : Schemeᵒᵖ} (f : X ⟶ Y) : Γ.map f = f.unop.appTop := rfl -theorem Γ_map_op {X Y : Scheme} (f : X ⟶ Y) : Γ.map f.op = f.app ⊤ := +theorem Γ_map_op {X Y : Scheme} (f : X ⟶ Y) : Γ.map f.op = f.appTop := rfl /-- @@ -466,13 +483,13 @@ def ΓSpecIso : Γ(Spec R, ⊤) ≅ R := SpecΓIdentity.app R @[reassoc (attr := simp)] lemma ΓSpecIso_naturality {R S : CommRingCat.{u}} (f : R ⟶ S) : - (Spec.map f).app ⊤ ≫ (ΓSpecIso S).hom = (ΓSpecIso R).hom ≫ f := SpecΓIdentity.hom.naturality f + (Spec.map f).appTop ≫ (ΓSpecIso S).hom = (ΓSpecIso R).hom ≫ f := SpecΓIdentity.hom.naturality f -- The RHS is not necessarily simpler than the LHS, but this direction coincides with the simp -- direction of `NatTrans.naturality`. @[reassoc (attr := simp)] lemma ΓSpecIso_inv_naturality {R S : CommRingCat.{u}} (f : R ⟶ S) : - f ≫ (ΓSpecIso S).inv = (ΓSpecIso R).inv ≫ (Spec.map f).app ⊤ := SpecΓIdentity.inv.naturality f + f ≫ (ΓSpecIso S).inv = (ΓSpecIso R).inv ≫ (Spec.map f).appTop := SpecΓIdentity.inv.naturality f -- This is not marked simp to respect the abstraction lemma ΓSpecIso_inv : (ΓSpecIso R).inv = StructureSheaf.toOpen R ⊤ := rfl @@ -538,6 +555,10 @@ theorem preimage_basicOpen {X Y : Scheme.{u}} (f : X ⟶ Y) {U : Y.Opens} (r : f ⁻¹ᵁ (Y.basicOpen r) = X.basicOpen (f.app U r) := LocallyRingedSpace.preimage_basicOpen f.toLRSHom r +theorem preimage_basicOpen_top {X Y : Scheme.{u}} (f : X ⟶ Y) (r : Γ(Y, ⊤)) : + f ⁻¹ᵁ (Y.basicOpen r) = X.basicOpen (f.appTop r) := + preimage_basicOpen .. + lemma basicOpen_appLE {X Y : Scheme.{u}} (f : X ⟶ Y) (U : X.Opens) (V : Y.Opens) (e : U ≤ f ⁻¹ᵁ V) (s : Γ(Y, V)) : X.basicOpen (f.appLE V U e s) = U ⊓ f ⁻¹ᵁ (Y.basicOpen s) := by simp only [preimage_basicOpen, Hom.appLE, CommRingCat.coe_comp_of, RingHom.coe_comp, diff --git a/Mathlib/AlgebraicGeometry/Stalk.lean b/Mathlib/AlgebraicGeometry/Stalk.lean index 6f520c2514a8f..e06eb1b7ef0ca 100644 --- a/Mathlib/AlgebraicGeometry/Stalk.lean +++ b/Mathlib/AlgebraicGeometry/Stalk.lean @@ -112,6 +112,13 @@ lemma fromSpecStalk_app {x : X} (hxU : x ∈ U) : hV.fromSpec_app_of_le _ hVU, ← X.presheaf.germ_res (homOfLE hVU) x hxV] simp [Category.assoc, ← ΓSpecIso_inv_naturality_assoc] +lemma fromSpecStalk_appTop {x : X} : + (X.fromSpecStalk x).appTop = + X.presheaf.germ ⊤ x trivial ≫ + (ΓSpecIso (X.presheaf.stalk x)).inv ≫ + (Spec (X.presheaf.stalk x)).presheaf.map (homOfLE le_top).op := + fromSpecStalk_app .. + @[reassoc (attr := simp)] lemma Spec_map_stalkSpecializes_fromSpecStalk {x y : X} (h : x ⤳ y) : Spec.map (X.presheaf.stalkSpecializes h) ≫ X.fromSpecStalk y = X.fromSpecStalk x := by @@ -184,7 +191,7 @@ instance {X : Scheme.{u}} (U : X.Opens) (x : X) (hxU : x ∈ U) : lemma fromSpecStalk_toSpecΓ (X : Scheme.{u}) (x : X) : X.fromSpecStalk x ≫ X.toSpecΓ = Spec.map (X.presheaf.germ ⊤ x trivial) := by rw [Scheme.toSpecΓ_naturality, ← SpecMap_ΓSpecIso_hom, ← Spec.map_comp, - @Scheme.fromSpecStalk_app X ⊤ _ trivial] + Scheme.fromSpecStalk_appTop] simp @[reassoc (attr := simp)] diff --git a/Mathlib/AlgebraicGeometry/ValuativeCriterion.lean b/Mathlib/AlgebraicGeometry/ValuativeCriterion.lean index cb9082af37462..dbc4ac5a16007 100644 --- a/Mathlib/AlgebraicGeometry/ValuativeCriterion.lean +++ b/Mathlib/AlgebraicGeometry/ValuativeCriterion.lean @@ -275,7 +275,7 @@ lemma IsSeparated.valuativeCriterion [IsSeparated f] : ValuativeCriterion.Unique conv_lhs => rw [← pullback.lift_fst l₁ l₂ h, ← pullback.condition_assoc] conv_rhs => rw [← pullback.lift_snd l₁ l₂ h, ← pullback.condition_assoc] simp - suffices h : Function.Bijective (g.app ⊤) by + suffices h : Function.Bijective (g.appTop) by refine (HasAffineProperty.iff_of_isAffine (P := MorphismProperty.isomorphisms Scheme)).mpr ?_ exact ⟨hZ, (ConcreteCategory.isIso_iff_bijective _).mpr h⟩ constructor @@ -286,7 +286,7 @@ lemma IsSeparated.valuativeCriterion [IsSeparated f] : ValuativeCriterion.Unique simpa using hl₂.symm have hg : l ≫ g = Spec.map (CommRingCat.ofHom (algebraMap S.R S.K)) := pullback.lift_snd _ _ _ - have : Function.Injective ((l ≫ g).app ⊤) := by + have : Function.Injective ((l ≫ g).appTop) := by rw [hg] let e := arrowIsoΓSpecOfIsAffine (CommRingCat.ofHom <| algebraMap S.R S.K) let P : MorphismProperty CommRingCat := @@ -296,7 +296,7 @@ lemma IsSeparated.valuativeCriterion [IsSeparated f] : ValuativeCriterion.Unique show P _ rw [← MorphismProperty.arrow_mk_iso_iff (P := P) e] exact NoZeroSMulDivisors.algebraMap_injective S.R S.K - rw [Scheme.comp_app _ _] at this + rw [Scheme.comp_appTop] at this exact Function.Injective.of_comp this · rw [@HasAffineProperty.iff_of_isAffine @IsClosedImmersion] at this exact this.right From f28fa7213f2fe9065a3a12f88d256ea7dd006d7e Mon Sep 17 00:00:00 2001 From: Pieter Cuijpers Date: Sun, 24 Nov 2024 20:04:13 +0000 Subject: [PATCH 259/829] feat: introduce the missing notation for `AddHom` (#19387) I noticed that there is notation for a non-unital monoid hom but not for a non-unital additive monoid hom, and since I'm working on extensions of non-unital homomorphisms in my PR's on quantale, I found it natural to add this. --- Mathlib/Algebra/Group/Hom/Defs.lean | 8 +++++++- 1 file changed, 7 insertions(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Group/Hom/Defs.lean b/Mathlib/Algebra/Group/Hom/Defs.lean index f2540ca0f00ab..ea061a0993f08 100644 --- a/Mathlib/Algebra/Group/Hom/Defs.lean +++ b/Mathlib/Algebra/Group/Hom/Defs.lean @@ -30,6 +30,7 @@ building blocks for other homomorphisms: * `→+`: Bundled `AddMonoid` homs. Also use for `AddGroup` homs. * `→*`: Bundled `Monoid` homs. Also use for `Group` homs. +* `→ₙ+`: Bundled `AddSemigroup` homs. * `→ₙ*`: Bundled `Semigroup` homs. ## Implementation notes @@ -95,7 +96,9 @@ end Zero section Add -/-- `AddHom M N` is the type of functions `M → N` that preserve addition. +/-- `M →ₙ+ N` is the type of functions `M → N` that preserve addition. The `ₙ` in the notation +stands for "non-unital" because it is intended to match the notation for `NonUnitalAlgHom` and +`NonUnitalRingHom`, so a `AddHom` is a non-unital additive monoid hom. When possible, instead of parametrizing results over `(f : AddHom M N)`, you should parametrize over `(F : Type*) [AddHomClass F M N] (f : F)`. @@ -108,6 +111,9 @@ structure AddHom (M : Type*) (N : Type*) [Add M] [Add N] where /-- The proposition that the function preserves addition -/ protected map_add' : ∀ x y, toFun (x + y) = toFun x + toFun y +/-- `M →ₙ+ N` denotes the type of addition-preserving maps from `M` to `N`. -/ +infixr:25 " →ₙ+ " => AddHom + /-- `AddHomClass F M N` states that `F` is a type of addition-preserving homomorphisms. You should declare an instance of this typeclass when you extend `AddHom`. -/ From 5a054eff5942f119027ee221b975c4488dda9ed3 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Mon, 25 Nov 2024 01:06:41 +0000 Subject: [PATCH 260/829] chore: update Mathlib dependencies 2024-11-25 (#19446) 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 6e18dc8970e13..aa466b92f0770 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "0dc51ac7947ff6aa2c16bcffb64c46c7149d1276", + "rev": "44e2d2e643fd2618b01f9a0592d7dcbd3ffa22de", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From 386ab515e2f0ed7be0211cbf37b1cd40b8cb4321 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 25 Nov 2024 03:22:16 +0000 Subject: [PATCH 261/829] chore(Multilinear/Basic): address a porting note (#19443) --- Mathlib/LinearAlgebra/Multilinear/Basic.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index 96de9639e3cfd..d12337d2b24a3 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -285,7 +285,6 @@ def constOfIsEmpty [IsEmpty ι] (m : M₂) : MultilinearMap R M₁ M₂ where end --- Porting note: Included `DFunLike.coe` to avoid strange CoeFun instance for Equiv /-- Given a multilinear map `f` on `n` variables (parameterized by `Fin n`) and a subset `s` of `k` of these variables, one gets a new multilinear map on `Fin k` by varying these variables, and fixing the other ones equal to a given value `z`. It is denoted by `f.restr s hk z`, where `hk` is a @@ -293,7 +292,7 @@ proof that the cardinality of `s` is `k`. The implicit identification between `F we use is the canonical (increasing) bijection. -/ def restr {k n : ℕ} (f : MultilinearMap R (fun _ : Fin n => M') M₂) (s : Finset (Fin n)) (hk : #s = k) (z : M') : MultilinearMap R (fun _ : Fin k => M') M₂ where - toFun v := f fun j => if h : j ∈ s then v ((DFunLike.coe (s.orderIsoOfFin hk).symm) ⟨j, h⟩) else z + toFun v := f fun j => if h : j ∈ s then v ((s.orderIsoOfFin hk).symm ⟨j, h⟩) else z /- Porting note: The proofs of the following two lemmas used to only use `erw` followed by `simp`, but it seems `erw` no longer unfolds or unifies well enough to work without more help. -/ map_update_add' v i x y := by From 04044ff3f96e5c4a520922b00450ef0dbef6e4e8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Nov 2024 03:43:22 +0000 Subject: [PATCH 262/829] chore: move pointwise set files to `Algebra._.Pointwise` (#19002) --- Mathlib.lean | 8 ++++---- Mathlib/Algebra/Algebra/Operations.lean | 2 +- Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean | 4 ++-- Mathlib/Algebra/Group/Pointwise/Finset/Interval.lean | 2 +- .../Group/Pointwise/Set}/BigOperators.lean | 0 Mathlib/Algebra/Group/Pointwise/Set/Card.lean | 2 +- .../Pointwise => Algebra/Group/Pointwise/Set}/Finite.lean | 0 .../Group/Pointwise/Set}/ListOfFn.lean | 0 .../Set => Algebra/Order/Group}/Pointwise/Interval.lean | 0 Mathlib/Analysis/Normed/Field/Basic.lean | 4 ++-- Mathlib/Data/Real/Cardinality.lean | 2 +- Mathlib/GroupTheory/Finiteness.lean | 2 +- Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean | 2 +- Mathlib/Topology/Algebra/Field.lean | 2 +- Mathlib/Topology/Algebra/Order/Field.lean | 2 +- Mathlib/Topology/MetricSpace/Pseudo/Real.lean | 2 +- 16 files changed, 17 insertions(+), 17 deletions(-) rename Mathlib/{Data/Set/Pointwise => Algebra/Group/Pointwise/Set}/BigOperators.lean (100%) rename Mathlib/{Data/Set/Pointwise => Algebra/Group/Pointwise/Set}/Finite.lean (100%) rename Mathlib/{Data/Set/Pointwise => Algebra/Group/Pointwise/Set}/ListOfFn.lean (100%) rename Mathlib/{Data/Set => Algebra/Order/Group}/Pointwise/Interval.lean (100%) diff --git a/Mathlib.lean b/Mathlib.lean index 58d4cfdeb00c9..f3b4520b86d56 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -294,7 +294,10 @@ 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.Set.Basic +import Mathlib.Algebra.Group.Pointwise.Set.BigOperators import Mathlib.Algebra.Group.Pointwise.Set.Card +import Mathlib.Algebra.Group.Pointwise.Set.Finite +import Mathlib.Algebra.Group.Pointwise.Set.ListOfFn import Mathlib.Algebra.Group.Prod import Mathlib.Algebra.Group.Semiconj.Basic import Mathlib.Algebra.Group.Semiconj.Defs @@ -661,6 +664,7 @@ import Mathlib.Algebra.Order.Group.OrderIso import Mathlib.Algebra.Order.Group.PiLex import Mathlib.Algebra.Order.Group.Pointwise.Bounds import Mathlib.Algebra.Order.Group.Pointwise.CompleteLattice +import Mathlib.Algebra.Order.Group.Pointwise.Interval import Mathlib.Algebra.Order.Group.PosPart import Mathlib.Algebra.Order.Group.Prod import Mathlib.Algebra.Order.Group.Synonym @@ -2820,11 +2824,7 @@ import Mathlib.Data.Set.Operations import Mathlib.Data.Set.Opposite import Mathlib.Data.Set.Pairwise.Basic import Mathlib.Data.Set.Pairwise.Lattice -import Mathlib.Data.Set.Pointwise.BigOperators -import Mathlib.Data.Set.Pointwise.Finite -import Mathlib.Data.Set.Pointwise.Interval import Mathlib.Data.Set.Pointwise.Iterate -import Mathlib.Data.Set.Pointwise.ListOfFn import Mathlib.Data.Set.Pointwise.SMul import Mathlib.Data.Set.Pointwise.Support import Mathlib.Data.Set.Prod diff --git a/Mathlib/Algebra/Algebra/Operations.lean b/Mathlib/Algebra/Algebra/Operations.lean index de1a4621d5777..960f805fd05b7 100644 --- a/Mathlib/Algebra/Algebra/Operations.lean +++ b/Mathlib/Algebra/Algebra/Operations.lean @@ -6,9 +6,9 @@ Authors: Kenny Lau import Mathlib.Algebra.Algebra.Bilinear import Mathlib.Algebra.Algebra.Opposite import Mathlib.Algebra.Group.Pointwise.Finset.Basic +import Mathlib.Algebra.Group.Pointwise.Set.BigOperators import Mathlib.Algebra.GroupWithZero.NonZeroDivisors import Mathlib.Algebra.Module.Submodule.Pointwise -import Mathlib.Data.Set.Pointwise.BigOperators import Mathlib.Data.Set.Semiring import Mathlib.GroupTheory.GroupAction.SubMulAction.Pointwise diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index a7cd803e05c07..4330e3cf9abae 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -5,11 +5,11 @@ Authors: Floris van Doorn, Yaël Dillies -/ import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Algebra.Group.Action.Pi +import Mathlib.Algebra.Group.Pointwise.Set.Finite +import Mathlib.Algebra.Group.Pointwise.Set.ListOfFn import Mathlib.Data.Finset.Density import Mathlib.Data.Finset.Max import Mathlib.Data.Finset.NAry -import Mathlib.Data.Set.Pointwise.Finite -import Mathlib.Data.Set.Pointwise.ListOfFn import Mathlib.Data.Set.Pointwise.SMul /-! diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Interval.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Interval.lean index 479121a3d4a1b..d06c3e7259b13 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Interval.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Interval.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ import Mathlib.Algebra.Group.Pointwise.Finset.Basic -import Mathlib.Data.Set.Pointwise.Interval +import Mathlib.Algebra.Order.Group.Pointwise.Interval import Mathlib.Order.Interval.Finset.Defs /-! # Pointwise operations on intervals diff --git a/Mathlib/Data/Set/Pointwise/BigOperators.lean b/Mathlib/Algebra/Group/Pointwise/Set/BigOperators.lean similarity index 100% rename from Mathlib/Data/Set/Pointwise/BigOperators.lean rename to Mathlib/Algebra/Group/Pointwise/Set/BigOperators.lean diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean index 06a3718af24d0..9472500e7fcc9 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean @@ -3,8 +3,8 @@ Copyright (c) 2024 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ +import Mathlib.Algebra.Group.Pointwise.Set.Finite import Mathlib.Data.Set.Card -import Mathlib.Data.Set.Pointwise.Finite import Mathlib.SetTheory.Cardinal.Finite /-! diff --git a/Mathlib/Data/Set/Pointwise/Finite.lean b/Mathlib/Algebra/Group/Pointwise/Set/Finite.lean similarity index 100% rename from Mathlib/Data/Set/Pointwise/Finite.lean rename to Mathlib/Algebra/Group/Pointwise/Set/Finite.lean diff --git a/Mathlib/Data/Set/Pointwise/ListOfFn.lean b/Mathlib/Algebra/Group/Pointwise/Set/ListOfFn.lean similarity index 100% rename from Mathlib/Data/Set/Pointwise/ListOfFn.lean rename to Mathlib/Algebra/Group/Pointwise/Set/ListOfFn.lean diff --git a/Mathlib/Data/Set/Pointwise/Interval.lean b/Mathlib/Algebra/Order/Group/Pointwise/Interval.lean similarity index 100% rename from Mathlib/Data/Set/Pointwise/Interval.lean rename to Mathlib/Algebra/Order/Group/Pointwise/Interval.lean diff --git a/Mathlib/Analysis/Normed/Field/Basic.lean b/Mathlib/Analysis/Normed/Field/Basic.lean index a81ab41393a9d..eb3c22b61051f 100644 --- a/Mathlib/Analysis/Normed/Field/Basic.lean +++ b/Mathlib/Analysis/Normed/Field/Basic.lean @@ -5,10 +5,10 @@ Authors: Patrick Massot, Johannes Hölzl -/ import Mathlib.Algebra.Algebra.NonUnitalSubalgebra import Mathlib.Algebra.Algebra.Subalgebra.Basic +import Mathlib.Algebra.Field.Subfield.Defs +import Mathlib.Algebra.Order.Group.Pointwise.Interval import Mathlib.Analysis.Normed.Group.Constructions import Mathlib.Analysis.Normed.Group.Submodule -import Mathlib.Data.Set.Pointwise.Interval -import Mathlib.Algebra.Field.Subfield.Defs /-! # Normed fields diff --git a/Mathlib/Data/Real/Cardinality.lean b/Mathlib/Data/Real/Cardinality.lean index b0a7e99a26fb2..b41af8e41d56f 100644 --- a/Mathlib/Data/Real/Cardinality.lean +++ b/Mathlib/Data/Real/Cardinality.lean @@ -3,9 +3,9 @@ Copyright (c) 2019 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.Order.Group.Pointwise.Interval import Mathlib.Analysis.SpecificLimits.Basic import Mathlib.Data.Rat.Cardinal -import Mathlib.Data.Set.Pointwise.Interval import Mathlib.SetTheory.Cardinal.Continuum /-! diff --git a/Mathlib/GroupTheory/Finiteness.lean b/Mathlib/GroupTheory/Finiteness.lean index 80895c9fc92ee..bd31d6b8743b4 100644 --- a/Mathlib/GroupTheory/Finiteness.lean +++ b/Mathlib/GroupTheory/Finiteness.lean @@ -3,8 +3,8 @@ Copyright (c) 2021 Riccardo Brasca. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Riccardo Brasca -/ +import Mathlib.Algebra.Group.Pointwise.Set.Finite import Mathlib.Algebra.Group.Subgroup.Pointwise -import Mathlib.Data.Set.Pointwise.Finite import Mathlib.GroupTheory.QuotientGroup.Defs import Mathlib.SetTheory.Cardinal.Finite diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean index 71f99e5c85dde..259da107ba011 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Joseph Myers. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers -/ -import Mathlib.Data.Set.Pointwise.Interval +import Mathlib.Algebra.Order.Group.Pointwise.Interval import Mathlib.LinearAlgebra.AffineSpace.Basic import Mathlib.LinearAlgebra.BilinearMap import Mathlib.LinearAlgebra.Pi diff --git a/Mathlib/Topology/Algebra/Field.lean b/Mathlib/Topology/Algebra/Field.lean index 7e78ef983d7f6..cc34c6223561e 100644 --- a/Mathlib/Topology/Algebra/Field.lean +++ b/Mathlib/Topology/Algebra/Field.lean @@ -5,10 +5,10 @@ Authors: Patrick Massot, Kim Morrison -/ import Mathlib.Algebra.Field.Subfield.Basic import Mathlib.Algebra.GroupWithZero.Divisibility +import Mathlib.Algebra.Order.Group.Pointwise.Interval import Mathlib.Topology.Algebra.GroupWithZero import Mathlib.Topology.Algebra.Ring.Basic import Mathlib.Topology.Order.LocalExtr -import Mathlib.Data.Set.Pointwise.Interval /-! # Topological fields diff --git a/Mathlib/Topology/Algebra/Order/Field.lean b/Mathlib/Topology/Algebra/Order/Field.lean index 51f7585e7e00c..ee7ea3caa6442 100644 --- a/Mathlib/Topology/Algebra/Order/Field.lean +++ b/Mathlib/Topology/Algebra/Order/Field.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Benjamin Davidson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Benjamin Davidson, Devon Tuma, Eric Rodriguez, Oliver Nash -/ -import Mathlib.Data.Set.Pointwise.Interval +import Mathlib.Algebra.Order.Group.Pointwise.Interval import Mathlib.Order.Filter.AtTopBot.Field import Mathlib.Topology.Algebra.Field import Mathlib.Topology.Algebra.Order.Group diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Real.lean b/Mathlib/Topology/MetricSpace/Pseudo/Real.lean index ae477178577c5..baa4f534f75cc 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Real.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Real.lean @@ -3,7 +3,7 @@ Copyright (c) 2015, 2017 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Robert Y. Lewis, Johannes Hölzl, Mario Carneiro, Sébastien Gouëzel -/ -import Mathlib.Data.Set.Pointwise.Interval +import Mathlib.Algebra.Order.Group.Pointwise.Interval import Mathlib.Topology.MetricSpace.Pseudo.Pi /-! From aa050b47ffec260eb307d5ad30c2cc02a0a8c5d4 Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 25 Nov 2024 06:35:02 +0000 Subject: [PATCH 263/829] feat(bench_summary): report no significant changes and job_id (#18812) If the `!bench` summary contains no entry with a difference in instructions of at least 10^9, then the bot posts a message confirming this information, rather than silently failing. The bot also posts a link to the workflow CI run that generated the message. --- .github/workflows/bench_summary_comment.yml | 2 +- scripts/bench_summary.lean | 16 ++++++++++++---- 2 files changed, 13 insertions(+), 5 deletions(-) diff --git a/.github/workflows/bench_summary_comment.yml b/.github/workflows/bench_summary_comment.yml index 6672fb8443e98..fe14ecdc32a79 100644 --- a/.github/workflows/bench_summary_comment.yml +++ b/.github/workflows/bench_summary_comment.yml @@ -27,7 +27,7 @@ jobs: run: | { cat scripts/bench_summary.lean - printf $'run_cmd BenchAction.addBenchSummaryComment %s "leanprover-community/mathlib4"' "${PR}" + printf $'run_cmd BenchAction.addBenchSummaryComment %s "leanprover-community/mathlib4" %s' "${PR}" "${{ github.run_id }}" } | lake env lean --stdin env: diff --git a/scripts/bench_summary.lean b/scripts/bench_summary.lean index 9c9407ea066f4..17f0467b2ef97 100644 --- a/scripts/bench_summary.lean +++ b/scripts/bench_summary.lean @@ -147,6 +147,7 @@ open Lean Elab Command in as a comment to a pull request. It takes as input * the number `PR` and the name `repo` as a `String` containing the relevant pull-request (it reads and posts comments there) +* the optional `jobID` string for reporting the action that produced the output * the `String` `tempFile` of a temporary file where the command stores transient information. The code itself interfaces with the shell to retrieve and process json data and eventually @@ -162,9 +163,10 @@ Here is a summary of the steps: * process the final string to produce a summary (using `benchOutput`), * finally post the resulting output to the PR (using `gh pr comment ...`). -/ -def addBenchSummaryComment (PR : Nat) (repo : String) +def addBenchSummaryComment (PR : Nat) (repo : String) (jobID : String := "") (author : String := "leanprover-bot") (tempFile : String := "benchOutput.json") : CommandElabM Unit := do + let job_msg := s!"\n[CI run](https://github.com/{repo}/actions/runs/{jobID})" let PR := s!"{PR}" let jq := s!".comments | last | select(.author.login==\"{author}\") | .body" @@ -197,10 +199,10 @@ def addBenchSummaryComment (PR : Nat) (repo : String) IO.FS.writeFile (tempFile ++ ".src") bench -- Extract all instruction changes whose magnitude is larger than `threshold`. - let threshold := s!"{10 ^ 9}" + let threshold := 10 ^ 9 let jq1 : IO.Process.SpawnArgs := { cmd := "jq" - args := #["-r", "--arg", "thr", threshold, + args := #["-r", "--arg", "thr", s!"{threshold}", ".differences | .[] | ($thr|tonumber) as $th | select(.dimension.metric == \"instructions\" and ((.diff >= $th) or (.diff <= -$th)))", (tempFile ++ ".src")] } @@ -220,6 +222,12 @@ def addBenchSummaryComment (PR : Nat) (repo : String) jq -c '[\{file: .dimension.benchmark, diff: .diff, reldiff: .reldiff}]' {tempFile} > \ {tempFile}.2" let secondFilter ← IO.Process.run jq2 + if secondFilter == "" then + let _ ← IO.Process.run + { cmd := "gh", args := #["pr", "comment", PR, "--repo", repo, "--body", + s!"No benchmark entry differed by at least {formatDiff threshold} instructions." ++ + job_msg] } + else IO.FS.writeFile tempFile secondFilter let jq3 : IO.Process.SpawnArgs := { cmd := "jq", args := #["-n", "reduce inputs as $in (null; . + $in)", tempFile] } @@ -230,7 +238,7 @@ def addBenchSummaryComment (PR : Nat) (repo : String) IO.println report -- Post the computed summary as a github comment. let add_comment : IO.Process.SpawnArgs := - { cmd := "gh", args := #["pr", "comment", PR, "--repo", repo, "--body", report] } + { cmd := "gh", args := #["pr", "comment", PR, "--repo", repo, "--body", report ++ job_msg] } let _ ← IO.Process.run add_comment end BenchAction From aad267d3a84144ad19fe4ea77f9ad0fd80679f95 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Mon, 25 Nov 2024 07:23:24 +0000 Subject: [PATCH 264/829] chore(discover-lean-pr-testing): ignore stage0 PRs, more tracing (#19405) --- .github/workflows/discover-lean-pr-testing.yml | 18 +++++++++++++----- 1 file changed, 13 insertions(+), 5 deletions(-) diff --git a/.github/workflows/discover-lean-pr-testing.yml b/.github/workflows/discover-lean-pr-testing.yml index 8210a3e66daab..f9df1791979d7 100644 --- a/.github/workflows/discover-lean-pr-testing.yml +++ b/.github/workflows/discover-lean-pr-testing.yml @@ -36,8 +36,10 @@ jobs: # The "./" in front of the path is important for `git show` OLD=$(git show "$PREV_COMMIT":./lean-toolchain | cut -f2 -d:) - echo "new=$NEW" >> "$GITHUB_OUTPUT" + echo "old=$OLD" + echo "new=$NEW" echo "old=$OLD" >> "$GITHUB_OUTPUT" + echo "new=$NEW" >> "$GITHUB_OUTPUT" - name: Clone lean4 repository and get PRs id: get-prs @@ -63,7 +65,10 @@ jobs: # This will only fetch commits newer than previously fetched commits (ie $OLD) git fetch origin tag "$NEW" --no-tags - PRS=$(git log --oneline "$OLD..$NEW" | sed 's/.*(#\([0-9]\+\))$/\1/') + # Find all commits to lean4 between the $OLD and $NEW toolchains + # and extract the github PR numbers + # (drop anything that doesn't look like a PR number from the final result) + PRS=$(git log --oneline "$OLD..$NEW" | sed 's/.*(#\([0-9]\+\))$/\1/' | grep -E '^[0-9]+$') # Output the PRs echo "$PRS" @@ -74,12 +79,14 @@ jobs: run: | # Use the PRs from the previous step PRS="${{ steps.get-prs.outputs.prs }}" + echo "=== $PRS ========================" echo "$PRS" - echo "====================" echo "$PRS" | tr ' ' '\n' > prs.txt + echo "=== prs.txt =====================" + cat prs.txt MATCHING_BRANCHES=$(git branch -r | grep -f prs.txt) + echo "=== $MATCHING_BRANCHES ==========" echo "$MATCHING_BRANCHES" - echo "====================" # Initialize an empty variable to store branches with relevant diffs RELEVANT_BRANCHES="" @@ -92,11 +99,12 @@ jobs: # Check if the diff contains files other than the specified ones if echo "$DIFF_FILES" | grep -v -e 'lake-manifest.json' -e 'lakefile.lean' -e 'lean-toolchain'; then # If it does, add the branch to RELEVANT_BRANCHES - RELEVANT_BRANCHES="$RELEVANT_BRANCHES\n- $BRANCH" + RELEVANT_BRANCHES="$RELEVANT_BRANCHES"$'\n- '"$BRANCH" fi done # Output the relevant branches + echo "=== $RELEVANT_BRANCHES ==========" echo "'$RELEVANT_BRANCHES'" printf "branches<> "$GITHUB_OUTPUT" From 55a97c9bd9fc11ce7a0cec05b9b355ffcbaa63ba Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Mon, 25 Nov 2024 07:56:44 +0000 Subject: [PATCH 265/829] chore(discover-lean-pr-testing): fix escaping hazard (#19448) --- .github/workflows/discover-lean-pr-testing.yml | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/.github/workflows/discover-lean-pr-testing.yml b/.github/workflows/discover-lean-pr-testing.yml index f9df1791979d7..02e087dfb6085 100644 --- a/.github/workflows/discover-lean-pr-testing.yml +++ b/.github/workflows/discover-lean-pr-testing.yml @@ -79,13 +79,13 @@ jobs: run: | # Use the PRs from the previous step PRS="${{ steps.get-prs.outputs.prs }}" - echo "=== $PRS ========================" + echo "=== PRS =========================" echo "$PRS" echo "$PRS" | tr ' ' '\n' > prs.txt echo "=== prs.txt =====================" cat prs.txt MATCHING_BRANCHES=$(git branch -r | grep -f prs.txt) - echo "=== $MATCHING_BRANCHES ==========" + echo "=== MATCHING_BRANCHES ===========" echo "$MATCHING_BRANCHES" # Initialize an empty variable to store branches with relevant diffs @@ -104,7 +104,7 @@ jobs: done # Output the relevant branches - echo "=== $RELEVANT_BRANCHES ==========" + echo "=== RELEVANT_BRANCHES ===========" echo "'$RELEVANT_BRANCHES'" printf "branches<> "$GITHUB_OUTPUT" From a41d8f43ee87309b9bb3c4f23c473a5fcbc2e922 Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 25 Nov 2024 08:25:35 +0000 Subject: [PATCH 266/829] feat: trigger automated Zulip emojis via `bors x` command (#19371) Changes the way in which emoji reactions to `delegated`, `ready-to-merge` and `merged` are applied in Zulip: instead of running periodically a script to add the reactions, the reactions are now applied at the same time that the label is applied, or when PRs are merged into master. --- .github/workflows/maintainer_bors.yml | 28 ++++- .../workflows/zulip_emoji_merge_delegate.yaml | 21 ++-- scripts/README.md | 9 +- scripts/zulip_emoji_merge_delegate.py | 105 +++++++++++------- 4 files changed, 109 insertions(+), 54 deletions(-) diff --git a/.github/workflows/maintainer_bors.yml b/.github/workflows/maintainer_bors.yml index 8b06a44a1b926..290ea7c66fa1b 100644 --- a/.github/workflows/maintainer_bors.yml +++ b/.github/workflows/maintainer_bors.yml @@ -84,4 +84,30 @@ jobs: curl --request DELETE \ --url "https://api.github.com/repos/${{ github.repository }}/issues/${{ github.event.issue.number }}${{ github.event.pull_request.number }}/labels/awaiting-author" \ --header 'authorization: Bearer ${{ secrets.TRIAGE_TOKEN }}' - # comment uses ${{ secrets.GITHUB_TOKEN }} + + - name: Set up Python + if: ${{ ! steps.merge_or_delegate.outputs.mOrD == '' && + ( steps.user_permission.outputs.require-result == 'true' || + steps.merge_or_delegate.outputs.bot == 'true' ) }} + uses: actions/setup-python@v5 + with: + python-version: '3.x' + + - name: Install dependencies + if: ${{ ! steps.merge_or_delegate.outputs.mOrD == '' && + ( steps.user_permission.outputs.require-result == 'true' || + steps.merge_or_delegate.outputs.bot == 'true' ) }} + run: | + python -m pip install --upgrade pip + pip install zulip + + - name: Run Zulip Emoji Merge Delegate Script + if: ${{ ! steps.merge_or_delegate.outputs.mOrD == '' && + ( steps.user_permission.outputs.require-result == 'true' || + steps.merge_or_delegate.outputs.bot == 'true' ) }} + env: + ZULIP_API_KEY: ${{ secrets.ZULIP_API_KEY }} + ZULIP_EMAIL: github-mathlib4-bot@leanprover.zulipchat.com + ZULIP_SITE: https://leanprover.zulipchat.com + run: | + python scripts/zulip_emoji_merge_delegate.py "$ZULIP_API_KEY" "$ZULIP_EMAIL" "$ZULIP_SITE" "${{ steps.merge_or_delegate.outputs.mOrD }}" "${{ github.event.issue.number }}${{ github.event.pull_request.number }}" diff --git a/.github/workflows/zulip_emoji_merge_delegate.yaml b/.github/workflows/zulip_emoji_merge_delegate.yaml index e44b4d07004d9..3de0bc2fd1aa4 100644 --- a/.github/workflows/zulip_emoji_merge_delegate.yaml +++ b/.github/workflows/zulip_emoji_merge_delegate.yaml @@ -1,19 +1,19 @@ name: Zulip Emoji Merge Delegate on: - schedule: - - cron: '0 * * * *' # Runs every hour + push: + branches: + - master jobs: - zulip-emoji-merge-delegate: + zulip-emoji-merged: runs-on: ubuntu-latest steps: - name: Checkout mathlib repository uses: actions/checkout@v4 with: - sparse-checkout: | - scripts + fetch-depth: 0 # donwload the full repository - name: Set up Python uses: actions/setup-python@v5 @@ -30,6 +30,13 @@ jobs: ZULIP_API_KEY: ${{ secrets.ZULIP_API_KEY }} ZULIP_EMAIL: github-mathlib4-bot@leanprover.zulipchat.com ZULIP_SITE: https://leanprover.zulipchat.com - GITHUB_TOKEN: ${{ secrets.GITHUB_TOKEN }} run: | - python scripts/zulip_emoji_merge_delegate.py "$ZULIP_API_KEY" "$ZULIP_EMAIL" "$ZULIP_SITE" "$GITHUB_TOKEN" + # scan the commits of the past 10 minutes, assuming that the timezone of the current machine + # is the same that performed the commit + git log --since="10 minutes ago" --pretty=oneline + + printf $'Scanning commits:\n%s\n\nContinuing\n' "$(git log --since="10 minutes ago" --pretty=oneline)" + while read -r pr_number; do + printf $'Running the python script with pr "%s"\n' "${pr_number}" + python scripts/zulip_emoji_merge_delegate.py "$ZULIP_API_KEY" "$ZULIP_EMAIL" "$ZULIP_SITE" "[Merged by Bors]" "${pr_number}" + done < <(git log --oneline -n 10 | awk -F# '{ gsub(/)$/, ""); print $NF}') diff --git a/scripts/README.md b/scripts/README.md index 44156380d5e62..3962da0193667 100644 --- a/scripts/README.md +++ b/scripts/README.md @@ -93,9 +93,12 @@ please do not add new entries to these files. PRs removing (the need for) entrie to the appropriate topic on zulip. - `count-trans-deps.py`, `import-graph-report.py` and `import_trans_difference.sh` produce various summaries of changes in transitive imports that the `PR_summary` message incorporates. -- `zulip_emoji_merge_delegate.py` is called every hour by a Github action cronjob. - It looks through the latest 200 zulip posts: if a message mentions a PR that is delegated, or sent to bors, or merged, - then this script will post an emoji reaction `:peace_sign:`, or `:bors:`, or `:merge:` respectively to the message. +- `zulip_emoji_merge_delegate.py` is called + * every time a `bors d`, `bors merge` or `bors r+` comment is added to a PR and + * on every push to `master`. + It looks through all zulip posts containing a reference to the relevant PR + (delegated, or sent to bors, or merged) and it will post an emoji reaction + `:peace_sign:`, or `:bors:`, or `:merge:` respectively to the message. - `late_importers.sh` is the main script used by the `latest_import.yml` action: it formats the `linter.minImports` output, summarizing the data in a table. See the module docs of `late_importers.sh` for further details. diff --git a/scripts/zulip_emoji_merge_delegate.py b/scripts/zulip_emoji_merge_delegate.py index d715df3bfbf9a..b279fbc6e1e11 100644 --- a/scripts/zulip_emoji_merge_delegate.py +++ b/scripts/zulip_emoji_merge_delegate.py @@ -1,17 +1,20 @@ #!/usr/bin/env python3 import sys import zulip -import requests import re # Usage: -# python scripts/zulip_emoji_merge_delegate.py $ZULIP_API_KEY $ZULIP_EMAIL $ZULIP_SITE $GITHUB_TOKEN +# python scripts/zulip_emoji_merge_delegate.py $ZULIP_API_KEY $ZULIP_EMAIL $ZULIP_SITE $LABEL $PR_NUMBER # See .github/workflows/zulip_emoji_merge_delegate.yaml for the meaning of these variables ZULIP_API_KEY = sys.argv[1] ZULIP_EMAIL = sys.argv[2] ZULIP_SITE = sys.argv[3] -GITHUB_TOKEN = sys.argv[4] +LABEL = sys.argv[4] +PR_NUMBER = sys.argv[5] + +print(f"LABEL: '{LABEL}'") +print(f"PR_NUMBER: '{PR_NUMBER}'") # Initialize Zulip client client = zulip.Client( @@ -22,58 +25,74 @@ # Fetch the last 200 messages response = client.get_messages({ - "anchor": "newest", - "num_before": 200, - "num_after": 0, - "narrow": [{"operator": "channel", "operand": "PR reviews"}], + "operator": "search", + "operand": f"https://github\.com/leanprover-community/mathlib4/pull/{PR_NUMBER}", }) messages = response['messages'] -pr_pattern = re.compile(r'https://github\.com/leanprover-community/mathlib4/pull/(\d+)') for message in messages: content = message['content'] - match = pr_pattern.search(content) - if match: - pr_number = match.group(1) - # Check for emoji reactions - reactions = message['reactions'] - has_peace_sign = any(reaction['emoji_name'] == 'peace_sign' for reaction in reactions) - has_bors = any(reaction['emoji_name'] == 'bors' for reaction in reactions) - has_merge = any(reaction['emoji_name'] == 'merge' for reaction in reactions) + print(f"matched: '{message}'") + # Check for emoji reactions + reactions = message['reactions'] + has_peace_sign = any(reaction['emoji_name'] == 'peace_sign' for reaction in reactions) + has_bors = any(reaction['emoji_name'] == 'bors' for reaction in reactions) + has_merge = any(reaction['emoji_name'] == 'merge' for reaction in reactions) + + pr_url = f"https://api.github.com/repos/leanprover-community/mathlib4/pulls/{PR_NUMBER}" + + print('Removing peace_sign') + result = client.remove_reaction({ + "message_id": message['id'], + "emoji_name": "peace_sign" + }) + print(f"result: '{result}'") + print('Removing bors') + result = client.remove_reaction({ + "message_id": message['id'], + "emoji_name": "bors" + }) + print(f"result: '{result}'") + + print('Removing merge') + result = client.remove_reaction({ + "message_id": message['id'], + "emoji_name": "merge" + }) + print(f"result: '{result}'") - pr_url = f"https://api.github.com/repos/leanprover-community/mathlib4/pulls/{pr_number}" - pr_response = requests.get(pr_url, headers={"Authorization": GITHUB_TOKEN}) - pr_data = pr_response.json() - labels = [label['name'] for label in pr_data['labels']] + if 'delegated' == LABEL: + print('adding delegated') - if 'delegated' in labels: - client.add_reaction({ + client.add_reaction({ + "message_id": message['id'], + "emoji_name": "peace_sign" + }) + elif 'ready-to-merge' == LABEL: + print('adding ready-to-merge') + if has_peace_sign: + client.remove_reaction({ "message_id": message['id'], "emoji_name": "peace_sign" }) - elif 'ready-to-merge' in labels: - if has_peace_sign: - client.remove_reaction({ - "message_id": message['id'], - "emoji_name": "peace_sign" - }) - client.add_reaction({ + client.add_reaction({ + "message_id": message['id'], + "emoji_name": "bors" + }) + elif LABEL.startswith("[Merged by Bors]"): + print('adding [Merged by Bors]') + if has_peace_sign: + client.remove_reaction({ "message_id": message['id'], - "emoji_name": "bors" + "emoji_name": "peace_sign" }) - elif pr_data['title'].startswith("[Merged by Bors]"): - if has_peace_sign: - client.remove_reaction({ - "message_id": message['id'], - "emoji_name": "peace_sign" - }) - if has_bors: - client.remove_reaction({ - "message_id": message['id'], - "emoji_name": "bors" - }) - client.add_reaction({ + if has_bors: + client.remove_reaction({ "message_id": message['id'], - "emoji_name": "merge" + "emoji_name": "bors" }) + client.add_reaction({ + "message_id": message['id'], + "emoji_name": "merge" + }) From 4eab2d4746acfc4e831ac0bac6d3440679c55f5c Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Mon, 25 Nov 2024 08:40:48 +0000 Subject: [PATCH 267/829] feat: the free locus of a (finitely presented) module (#18691) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib.lean | 1 + .../Algebra/Module/FinitePresentation.lean | 35 ++- Mathlib/Algebra/Module/FreeLocus.lean | 227 ++++++++++++++++++ .../Algebra/Module/LocalizedModule/Basic.lean | 47 +++- Mathlib/LinearAlgebra/FreeModule/Basic.lean | 21 ++ Mathlib/RingTheory/Finiteness/Basic.lean | 11 + .../LocalProperties/Projective.lean | 43 +++- 7 files changed, 379 insertions(+), 6 deletions(-) create mode 100644 Mathlib/Algebra/Module/FreeLocus.lean diff --git a/Mathlib.lean b/Mathlib.lean index f3b4520b86d56..7e32e8e68409c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -518,6 +518,7 @@ import Mathlib.Algebra.Module.Equiv.Basic import Mathlib.Algebra.Module.Equiv.Defs import Mathlib.Algebra.Module.Equiv.Opposite import Mathlib.Algebra.Module.FinitePresentation +import Mathlib.Algebra.Module.FreeLocus import Mathlib.Algebra.Module.GradedModule import Mathlib.Algebra.Module.Hom import Mathlib.Algebra.Module.Injective diff --git a/Mathlib/Algebra/Module/FinitePresentation.lean b/Mathlib/Algebra/Module/FinitePresentation.lean index a4d6e21f87576..553e1aee1ced2 100644 --- a/Mathlib/Algebra/Module/FinitePresentation.lean +++ b/Mathlib/Algebra/Module/FinitePresentation.lean @@ -4,10 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ import Mathlib.LinearAlgebra.FreeModule.Finite.Basic +import Mathlib.LinearAlgebra.TensorProduct.RightExactness import Mathlib.LinearAlgebra.Isomorphisms import Mathlib.RingTheory.Finiteness.Projective -import Mathlib.RingTheory.Localization.Module +import Mathlib.RingTheory.Finiteness.TensorProduct +import Mathlib.RingTheory.Localization.BaseChange import Mathlib.RingTheory.Noetherian.Defs + /-! # Finitely Presented Modules @@ -232,6 +235,36 @@ section CommRing variable {R M N N'} [CommRing R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] variable [AddCommGroup N'] [Module R N'] (S : Submonoid R) (f : N →ₗ[R] N') [IsLocalizedModule S f] +open TensorProduct in +instance {A} [CommRing A] [Algebra R A] [Module.FinitePresentation R M] : + Module.FinitePresentation A (A ⊗[R] M) := by + classical + obtain ⟨n, f, hf⟩ := Module.Finite.exists_fin' R M + have inst := Module.finitePresentation_of_projective A (A ⊗[R] (Fin n → R)) + apply Module.finitePresentation_of_surjective (f.baseChange A) + (LinearMap.lTensor_surjective A hf) + have : Function.Exact ((LinearMap.ker f).subtype.baseChange A) (f.baseChange A) := + lTensor_exact A f.exact_subtype_ker_map hf + rw [LinearMap.exact_iff] at this + rw [this, ← Submodule.map_top] + apply Submodule.FG.map + have : Module.Finite R (LinearMap.ker f) := + ⟨(Submodule.fg_top _).mpr (Module.FinitePresentation.fg_ker f hf)⟩ + exact Module.Finite.out (R := A) (M := A ⊗[R] LinearMap.ker f) + +open TensorProduct in +lemma FinitePresentation.of_isBaseChange + {A} [CommRing A] [Algebra R A] [Module A N] [IsScalarTower R A N] + (f : M →ₗ[R] N) (h : IsBaseChange A f) [Module.FinitePresentation R M] : + Module.FinitePresentation A N := + Module.finitePresentation_of_surjective + h.equiv.toLinearMap h.equiv.surjective (by simpa using Submodule.fg_bot) + +open TensorProduct in +instance (S : Submonoid R) [Module.FinitePresentation R M] : + Module.FinitePresentation (Localization S) (LocalizedModule S M) := + FinitePresentation.of_isBaseChange (LocalizedModule.mkLinearMap S M) + ((isLocalizedModule_iff_isBaseChange S _ _).mp inferInstance) lemma Module.FinitePresentation.exists_lift_of_isLocalizedModule [h : Module.FinitePresentation R M] (g : M →ₗ[R] N') : diff --git a/Mathlib/Algebra/Module/FreeLocus.lean b/Mathlib/Algebra/Module/FreeLocus.lean new file mode 100644 index 0000000000000..9f42786132f02 --- /dev/null +++ b/Mathlib/Algebra/Module/FreeLocus.lean @@ -0,0 +1,227 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.AlgebraicGeometry.PrimeSpectrum.Basic +import Mathlib.RingTheory.Flat.Stability +import Mathlib.RingTheory.LocalProperties.Projective +import Mathlib.RingTheory.LocalRing.Module +import Mathlib.RingTheory.Localization.Free +import Mathlib.RingTheory.Localization.LocalizationLocalization +import Mathlib.Topology.LocallyConstant.Basic + +/-! + +# The free locus of a module + +## Main definitions and results + +Let `M` be a finitely presented `R`-module. +- `Module.freeLocus`: The set of points `x` in `Spec R` such that `Mₓ` is free over `Rₓ`. +- `Module.freeLocus_eq_univ_iff`: + The free locus is the whole `Spec R` if and only if `M` is projective. +- `Module.basicOpen_subset_freeLocus_iff`: `D(f)` is contained in the free locus if and only if + `M_f` is projective over `R_f`. +- `Module.rankAtStalk`: The function `Spec R → ℕ` sending `x` to `rank_{Rₓ} Mₓ`. +- `Module.isLocallyConstant_rankAtStalk`: + If `M` is flat over `R`, then `rankAtStalk` is locally constant. + +-/ + +universe uR uM + +variable (R : Type uR) (M : Type uM) [CommRing R] [AddCommGroup M] [Module R M] + +namespace Module + +open PrimeSpectrum TensorProduct + +/-- The free locus of a module, i.e. the set of primes `p` such that `Mₚ` is free over `Rₚ`. -/ +def freeLocus : Set (PrimeSpectrum R) := + { p | Module.Free (Localization.AtPrime p.asIdeal) (LocalizedModule p.asIdeal.primeCompl M) } + +variable {R M} + +lemma mem_freeLocus {p} : p ∈ freeLocus R M ↔ + Module.Free (Localization.AtPrime p.asIdeal) (LocalizedModule p.asIdeal.primeCompl M) := + Iff.rfl + +attribute [local instance] RingHomInvPair.of_ringEquiv in +lemma mem_freeLocus_of_isLocalization (p : PrimeSpectrum R) + (Rₚ Mₚ) [CommRing Rₚ] [Algebra R Rₚ] [IsLocalization.AtPrime Rₚ p.asIdeal] + [AddCommGroup Mₚ] [Module R Mₚ] (f : M →ₗ[R] Mₚ) [IsLocalizedModule p.asIdeal.primeCompl f] + [Module Rₚ Mₚ] [IsScalarTower R Rₚ Mₚ] : + p ∈ freeLocus R M ↔ Module.Free Rₚ Mₚ := by + apply Module.Free.iff_of_ringEquiv (IsLocalization.algEquiv p.asIdeal.primeCompl + (Localization.AtPrime p.asIdeal) Rₚ).toRingEquiv + refine { __ := IsLocalizedModule.iso p.asIdeal.primeCompl f, map_smul' := ?_ } + intro r x + obtain ⟨r, s, rfl⟩ := IsLocalization.mk'_surjective p.asIdeal.primeCompl r + apply ((Module.End_isUnit_iff _).mp (IsLocalizedModule.map_units f s)).1 + simp only [AddHom.toFun_eq_coe, LinearMap.coe_toAddHom, LinearEquiv.coe_coe, + algebraMap_end_apply, AlgEquiv.toRingEquiv_eq_coe, + AlgEquiv.toRingEquiv_toRingHom, RingHom.coe_coe, IsLocalization.algEquiv_apply, + IsLocalization.map_id_mk'] + simp only [← map_smul, ← smul_assoc, IsLocalization.smul_mk'_self, algebraMap_smul, + IsLocalization.map_id_mk'] + +attribute [local instance] RingHomInvPair.of_ringEquiv in +lemma mem_freeLocus_iff_tensor (p : PrimeSpectrum R) + (Rₚ) [CommRing Rₚ] [Algebra R Rₚ] [IsLocalization.AtPrime Rₚ p.asIdeal] : + p ∈ freeLocus R M ↔ Module.Free Rₚ (Rₚ ⊗[R] M) := by + have := (isLocalizedModule_iff_isBaseChange p.asIdeal.primeCompl _ _).mpr + (TensorProduct.isBaseChange R M Rₚ) + exact mem_freeLocus_of_isLocalization p Rₚ (f := TensorProduct.mk R Rₚ M 1) + +lemma freeLocus_congr {M'} [AddCommGroup M'] [Module R M'] (e : M ≃ₗ[R] M') : + freeLocus R M = freeLocus R M' := by + ext p + exact mem_freeLocus_of_isLocalization _ _ _ + (LocalizedModule.mkLinearMap p.asIdeal.primeCompl M' ∘ₗ e.toLinearMap) + +open TensorProduct in +lemma comap_freeLocus_le {A} [CommRing A] [Algebra R A] : + comap (algebraMap R A) ⁻¹' freeLocus R M ≤ freeLocus A (A ⊗[R] M) := by + intro p hp + let Rₚ := Localization.AtPrime (comap (algebraMap R A) p).asIdeal + let Aₚ := Localization.AtPrime p.asIdeal + rw [Set.mem_preimage, mem_freeLocus_iff_tensor _ Rₚ] at hp + rw [mem_freeLocus_iff_tensor _ Aₚ] + letI : Algebra Rₚ Aₚ := (Localization.localRingHom + (comap (algebraMap R A) p).asIdeal p.asIdeal (algebraMap R A) rfl).toAlgebra + have : IsScalarTower R Rₚ Aₚ := IsScalarTower.of_algebraMap_eq' + (by simp [RingHom.algebraMap_toAlgebra, Localization.localRingHom, + ← IsScalarTower.algebraMap_eq]) + let e := AlgebraTensorModule.cancelBaseChange R Rₚ Aₚ Aₚ M ≪≫ₗ + (AlgebraTensorModule.cancelBaseChange R A Aₚ Aₚ M).symm + exact .of_equiv e + +lemma freeLocus_localization (S : Submonoid R) : + freeLocus (Localization S) (LocalizedModule S M) = + comap (algebraMap R _) ⁻¹' freeLocus R M := by + ext p + simp only [Set.mem_preimage] + let p' := p.asIdeal.comap (algebraMap R _) + have hp' : S ≤ p'.primeCompl := fun x hx H ↦ + p.isPrime.ne_top (Ideal.eq_top_of_isUnit_mem _ H (IsLocalization.map_units _ ⟨x, hx⟩)) + let Rₚ := Localization.AtPrime p' + let Mₚ := LocalizedModule p'.primeCompl M + letI : Algebra (Localization S) Rₚ := + IsLocalization.localizationAlgebraOfSubmonoidLe _ _ S p'.primeCompl hp' + have : IsScalarTower R (Localization S) Rₚ := + IsLocalization.localization_isScalarTower_of_submonoid_le .. + have : IsLocalization.AtPrime Rₚ p.asIdeal := by + have := IsLocalization.isLocalization_of_submonoid_le (Localization S) Rₚ _ _ hp' + apply IsLocalization.isLocalization_of_is_exists_mul_mem _ + (Submonoid.map (algebraMap R (Localization S)) p'.primeCompl) + · rintro _ ⟨x, hx, rfl⟩; exact hx + · rintro ⟨x, hx⟩ + obtain ⟨x, s, rfl⟩ := IsLocalization.mk'_surjective S x + refine ⟨algebraMap _ _ s.1, x, fun H ↦ hx ?_, by simp⟩ + rw [IsLocalization.mk'_eq_mul_mk'_one] + exact Ideal.mul_mem_right _ _ H + letI : Module (Localization S) Mₚ := Module.compHom Mₚ (algebraMap _ Rₚ) + have : IsScalarTower R (Localization S) Mₚ := + ⟨fun r r' m ↦ show algebraMap _ Rₚ (r • r') • m = _ by + simp [Algebra.smul_def, ← IsScalarTower.algebraMap_apply, mul_smul]; rfl⟩ + have : IsScalarTower (Localization S) Rₚ Mₚ := + ⟨fun r r' m ↦ show _ = algebraMap _ Rₚ r • _ by rw [← mul_smul, ← Algebra.smul_def]⟩ + let l := (IsLocalizedModule.liftOfLE _ _ hp' (LocalizedModule.mkLinearMap S M) + (LocalizedModule.mkLinearMap p'.primeCompl M)).extendScalarsOfIsLocalization S + (Localization S) + have : IsLocalizedModule p.asIdeal.primeCompl l := by + have : IsLocalizedModule p'.primeCompl (l.restrictScalars R) := + inferInstanceAs (IsLocalizedModule p'.primeCompl + (IsLocalizedModule.liftOfLE _ _ hp' (LocalizedModule.mkLinearMap S M) + (LocalizedModule.mkLinearMap p'.primeCompl M))) + have : IsLocalizedModule (Algebra.algebraMapSubmonoid (Localization S) p'.primeCompl) l := + IsLocalizedModule.of_restrictScalars p'.primeCompl .. + apply IsLocalizedModule.of_exists_mul_mem + (Algebra.algebraMapSubmonoid (Localization S) p'.primeCompl) + · rintro _ ⟨x, hx, rfl⟩; exact hx + · rintro ⟨x, hx⟩ + obtain ⟨x, s, rfl⟩ := IsLocalization.mk'_surjective S x + refine ⟨algebraMap _ _ s.1, x, fun H ↦ hx ?_, by simp⟩ + rw [IsLocalization.mk'_eq_mul_mk'_one] + exact Ideal.mul_mem_right _ _ H + rw [mem_freeLocus_of_isLocalization (R := Localization S) p Rₚ Mₚ l] + rfl + +lemma freeLocus_eq_univ_iff [Module.FinitePresentation R M] : + freeLocus R M = Set.univ ↔ Module.Projective R M := by + simp_rw [Set.eq_univ_iff_forall, mem_freeLocus] + exact ⟨fun H ↦ Module.projective_of_localization_maximal fun I hI ↦ + have := H ⟨I, hI.isPrime⟩; .of_free, fun H x ↦ Module.free_of_flat_of_isLocalRing⟩ + +lemma freeLocus_eq_univ [Module.FinitePresentation R M] [Module.Flat R M] : + freeLocus R M = Set.univ := by + simp_rw [Set.eq_univ_iff_forall, mem_freeLocus] + exact fun x ↦ Module.free_of_flat_of_isLocalRing + +lemma basicOpen_subset_freeLocus_iff [Module.FinitePresentation R M] {f : R} : + (basicOpen f : Set (PrimeSpectrum R)) ⊆ freeLocus R M ↔ + Module.Projective (Localization.Away f) (LocalizedModule (.powers f) M) := by + rw [← freeLocus_eq_univ_iff, freeLocus_localization, + Set.preimage_eq_univ_iff, localization_away_comap_range _ f] + +lemma isOpen_freeLocus [Module.FinitePresentation R M] : + IsOpen (freeLocus R M) := by + refine isOpen_iff_forall_mem_open.mpr fun x hx ↦ ?_ + have : Module.Free _ _ := hx + obtain ⟨r, hr, hr', _⟩ := Module.FinitePresentation.exists_free_localizedModule_powers + x.asIdeal.primeCompl (LocalizedModule.mkLinearMap x.asIdeal.primeCompl M) + (Localization.AtPrime x.asIdeal) + exact ⟨basicOpen r, basicOpen_subset_freeLocus_iff.mpr inferInstance, (basicOpen r).2, hr⟩ + +variable (M) in +/-- The rank of `M` at the stalk of `p` is the rank of `Mₚ` as a `Rₚ`-module. -/ +noncomputable +def rankAtStalk (p : PrimeSpectrum R) : ℕ := + Module.finrank (Localization.AtPrime p.asIdeal) (LocalizedModule p.asIdeal.primeCompl M) + +lemma isLocallyConstant_rankAtStalk_freeLocus [Module.FinitePresentation R M] : + IsLocallyConstant (fun x : freeLocus R M ↦ rankAtStalk M x.1) := by + refine (IsLocallyConstant.iff_exists_open _).mpr fun ⟨x, hx⟩ ↦ ?_ + have : Module.Free _ _ := hx + obtain ⟨f, hf, hf', hf''⟩ := Module.FinitePresentation.exists_free_localizedModule_powers + x.asIdeal.primeCompl (LocalizedModule.mkLinearMap x.asIdeal.primeCompl M) + (Localization.AtPrime x.asIdeal) + refine ⟨Subtype.val ⁻¹' basicOpen f, (basicOpen f).2.preimage continuous_subtype_val, hf, ?_⟩ + rintro ⟨p, hp''⟩ hp + let p' := Algebra.algebraMapSubmonoid (Localization (.powers f)) p.asIdeal.primeCompl + have hp' : Submonoid.powers f ≤ p.asIdeal.primeCompl := by + simpa [Submonoid.powers_le, Ideal.primeCompl] + let Rₚ := Localization.AtPrime p.asIdeal + let Mₚ := LocalizedModule p.asIdeal.primeCompl M + letI : Algebra (Localization.Away f) Rₚ := + IsLocalization.localizationAlgebraOfSubmonoidLe _ _ (.powers f) p.asIdeal.primeCompl hp' + have : IsScalarTower R (Localization.Away f) Rₚ := + IsLocalization.localization_isScalarTower_of_submonoid_le .. + letI : Module (Localization.Away f) Mₚ := Module.compHom Mₚ (algebraMap _ Rₚ) + have : IsScalarTower R (Localization.Away f) Mₚ := + ⟨fun r r' m ↦ show algebraMap _ Rₚ (r • r') • m = _ by + simp [Algebra.smul_def, ← IsScalarTower.algebraMap_apply, mul_smul]; rfl⟩ + have : IsScalarTower (Localization.Away f) Rₚ Mₚ := + ⟨fun r r' m ↦ show _ = algebraMap _ Rₚ r • _ by rw [← mul_smul, ← Algebra.smul_def]⟩ + let l := (IsLocalizedModule.liftOfLE _ _ hp' (LocalizedModule.mkLinearMap (.powers f) M) + (LocalizedModule.mkLinearMap p.asIdeal.primeCompl M)).extendScalarsOfIsLocalization (.powers f) + (Localization.Away f) + have : IsLocalization p' Rₚ := + IsLocalization.isLocalization_of_submonoid_le (Localization.Away f) Rₚ _ _ hp' + have : IsLocalizedModule p.asIdeal.primeCompl (l.restrictScalars R) := + inferInstanceAs (IsLocalizedModule p.asIdeal.primeCompl + ((IsLocalizedModule.liftOfLE _ _ hp' (LocalizedModule.mkLinearMap (.powers f) M) + (LocalizedModule.mkLinearMap p.asIdeal.primeCompl M)))) + have : IsLocalizedModule (Algebra.algebraMapSubmonoid _ p.asIdeal.primeCompl) l := + IsLocalizedModule.of_restrictScalars p.asIdeal.primeCompl .. + have := Module.finrank_of_isLocalizedModule_of_free Rₚ p' l + simp [rankAtStalk, this, hf''] + +lemma isLocallyConstant_rankAtStalk [Module.FinitePresentation R M] [Module.Flat R M] : + IsLocallyConstant (rankAtStalk (R := R) M) := by + let e : freeLocus R M ≃ₜ PrimeSpectrum R := + (Homeomorph.setCongr freeLocus_eq_univ).trans (Homeomorph.Set.univ (PrimeSpectrum R)) + convert isLocallyConstant_rankAtStalk_freeLocus.comp_continuous e.symm.continuous + +end Module diff --git a/Mathlib/Algebra/Module/LocalizedModule/Basic.lean b/Mathlib/Algebra/Module/LocalizedModule/Basic.lean index cd73b52dd3c65..efe62ee72caee 100644 --- a/Mathlib/Algebra/Module/LocalizedModule/Basic.lean +++ b/Mathlib/Algebra/Module/LocalizedModule/Basic.lean @@ -560,7 +560,7 @@ lemma IsLocalizedModule.eq_iff_exists [IsLocalizedModule S f] {x₁ x₂} : simp_rw [f.map_smul_of_tower, Submonoid.smul_def, ← Module.algebraMap_end_apply R R] at h exact ((Module.End_isUnit_iff _).mp <| map_units f c).1 h -theorem IsLocalizedModule.of_linearEquiv (e : M' ≃ₗ[R] M'') [hf : IsLocalizedModule S f] : +instance IsLocalizedModule.of_linearEquiv (e : M' ≃ₗ[R] M'') [hf : IsLocalizedModule S f] : IsLocalizedModule S (e ∘ₗ f : M →ₗ[R] M'') where map_units s := by rw [show algebraMap R (Module.End R M'') s = e ∘ₗ (algebraMap R (Module.End R M') s) ∘ₗ e.symm @@ -576,6 +576,18 @@ theorem IsLocalizedModule.of_linearEquiv (e : M' ≃ₗ[R] M'') [hf : IsLocalize EmbeddingLike.apply_eq_iff_eq] at h exact hf.exists_of_eq h +instance IsLocalizedModule.of_linearEquiv_right (e : M'' ≃ₗ[R] M) [hf : IsLocalizedModule S f] : + IsLocalizedModule S (f ∘ₗ e : M'' →ₗ[R] M') where + map_units s := hf.map_units s + surj' x := by + obtain ⟨⟨p, s⟩, h⟩ := hf.surj' x + exact ⟨⟨e.symm p, s⟩, by simpa using h⟩ + exists_of_eq h := by + simp_rw [LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply] at h + obtain ⟨c, hc⟩ := hf.exists_of_eq h + exact ⟨c, by simpa only [Submonoid.smul_def, map_smul, e.symm_apply_apply] + using congr(e.symm $hc)⟩ + variable (M) in lemma isLocalizedModule_id (R') [CommSemiring R'] [Algebra R R'] [IsLocalization S R'] [Module R' M] [IsScalarTower R R' M] : IsLocalizedModule S (.id : M →ₗ[R] M) where @@ -703,6 +715,39 @@ instance localizedModuleIsLocalizedModule : LocalizedModule.mk_cancel t] exists_of_eq eq1 := by simpa only [eq_comm, one_smul] using LocalizedModule.mk_eq.mp eq1 +lemma IsLocalizedModule.of_restrictScalars (S : Submonoid R) + {N : Type*} [AddCommGroup N] [Module R N] [Module A M] [Module A N] + [IsScalarTower R A M] [IsScalarTower R A N] + (f : M →ₗ[A] N) [IsLocalizedModule S (f.restrictScalars R)] : + IsLocalizedModule (Algebra.algebraMapSubmonoid A S) f where + map_units x := by + obtain ⟨_, x, hx, rfl⟩ := x + have := IsLocalizedModule.map_units (f.restrictScalars R) ⟨x, hx⟩ + simp only [← IsScalarTower.algebraMap_apply, Module.End_isUnit_iff] at this ⊢ + exact this + surj' y := by + obtain ⟨⟨x, t⟩, e⟩ := IsLocalizedModule.surj S (f.restrictScalars R) y + exact ⟨⟨x, ⟨_, t, t.2, rfl⟩⟩, by simpa [Submonoid.smul_def] using e⟩ + exists_of_eq {x₁ x₂} e := by + obtain ⟨c, hc⟩ := IsLocalizedModule.exists_of_eq (S := S) (f := f.restrictScalars R) e + refine ⟨⟨_, c, c.2, rfl⟩, by simpa [Submonoid.smul_def]⟩ + +lemma IsLocalizedModule.of_exists_mul_mem {N : Type*} [AddCommGroup N] [Module R N] + (S T : Submonoid R) (h : S ≤ T) (h' : ∀ x : T, ∃ m : R, m * x ∈ S) + (f : M →ₗ[R] N) [IsLocalizedModule S f] : + IsLocalizedModule T f where + map_units x := by + obtain ⟨m, mx⟩ := h' x + have := IsLocalizedModule.map_units f ⟨_, mx⟩ + rw [map_mul, (Algebra.commute_algebraMap_left _ _).isUnit_mul_iff] at this + exact this.2 + surj' y := by + obtain ⟨⟨x, t⟩, e⟩ := IsLocalizedModule.surj S f y + exact ⟨⟨x, ⟨t, h t.2⟩⟩, e⟩ + exists_of_eq {x₁ x₂} e := by + obtain ⟨c, hc⟩ := IsLocalizedModule.exists_of_eq (S := S) (f := f) e + exact ⟨⟨c, h c.2⟩, hc⟩ + namespace IsLocalizedModule variable [IsLocalizedModule S f] diff --git a/Mathlib/LinearAlgebra/FreeModule/Basic.lean b/Mathlib/LinearAlgebra/FreeModule/Basic.lean index a53c69bdf4ff3..99e69913af94b 100644 --- a/Mathlib/LinearAlgebra/FreeModule/Basic.lean +++ b/Mathlib/LinearAlgebra/FreeModule/Basic.lean @@ -118,6 +118,27 @@ theorem of_equiv' {P : Type v} [AddCommMonoid P] [Module R P] (_ : Module.Free R (e : P ≃ₗ[R] N) : Module.Free R N := of_equiv e +attribute [local instance] RingHomInvPair.of_ringEquiv in +lemma of_ringEquiv {R R' M M'} [Semiring R] [AddCommMonoid M] [Module R M] + [Semiring R'] [AddCommMonoid M'] [Module R' M'] + (e₁ : R ≃+* R') (e₂ : M ≃ₛₗ[RingHomClass.toRingHom e₁] M') [Module.Free R M] : + Module.Free R' M' := by + let I := Module.Free.ChooseBasisIndex R M + obtain ⟨e₃ : M ≃ₗ[R] I →₀ R⟩ := Module.Free.chooseBasis R M + let e : M' ≃+ (I →₀ R') := + (e₂.symm.trans e₃).toAddEquiv.trans (Finsupp.mapRange.addEquiv (α := I) e₁.toAddEquiv) + have he (x) : e x = Finsupp.mapRange.addEquiv (α := I) e₁.toAddEquiv (e₃ (e₂.symm x)) := rfl + let e' : M' ≃ₗ[R'] (I →₀ R') := + { __ := e, map_smul' := fun m x ↦ Finsupp.ext fun i ↦ by simp [he, map_smulₛₗ] } + exact of_basis (.ofRepr e') + +attribute [local instance] RingHomInvPair.of_ringEquiv in +lemma iff_of_ringEquiv {R R' M M'} [Semiring R] [AddCommMonoid M] [Module R M] + [Semiring R'] [AddCommMonoid M'] [Module R' M'] + (e₁ : R ≃+* R') (e₂ : M ≃ₛₗ[RingHomClass.toRingHom e₁] M') : + Module.Free R M ↔ Module.Free R' M' := + ⟨fun _ ↦ of_ringEquiv e₁ e₂, fun _ ↦ of_ringEquiv e₁.symm e₂.symm⟩ + variable (R M N) /-- The module structure provided by `Semiring.toModule` is free. -/ diff --git a/Mathlib/RingTheory/Finiteness/Basic.lean b/Mathlib/RingTheory/Finiteness/Basic.lean index f68cde85ea80c..73248997ca909 100644 --- a/Mathlib/RingTheory/Finiteness/Basic.lean +++ b/Mathlib/RingTheory/Finiteness/Basic.lean @@ -119,6 +119,17 @@ theorem fg_restrictScalars {R S M : Type*} [CommSemiring R] [Semiring S] [Algebr use X exact (Submodule.restrictScalars_span R S h (X : Set M)).symm +lemma FG.of_restrictScalars (R) {A M} [CommSemiring R] [Semiring A] [AddCommMonoid M] + [Algebra R A] [Module R M] [Module A M] [IsScalarTower R A M] (S : Submodule A M) + (hS : (S.restrictScalars R).FG) : S.FG := by + obtain ⟨s, e⟩ := hS + refine ⟨s, Submodule.restrictScalars_injective R _ _ (le_antisymm ?_ ?_)⟩ + · show Submodule.span A s ≤ S + have := Submodule.span_le.mp e.le + rwa [Submodule.span_le] + · rw [← e] + exact Submodule.span_le_restrictScalars _ _ _ + theorem FG.stabilizes_of_iSup_eq {M' : Submodule R M} (hM' : M'.FG) (N : ℕ →o Submodule R M) (H : iSup N = M') : ∃ n, M' = N n := by obtain ⟨S, hS⟩ := hM' diff --git a/Mathlib/RingTheory/LocalProperties/Projective.lean b/Mathlib/RingTheory/LocalProperties/Projective.lean index a2c4615a7942e..d6c0268663ed7 100644 --- a/Mathlib/RingTheory/LocalProperties/Projective.lean +++ b/Mathlib/RingTheory/LocalProperties/Projective.lean @@ -5,9 +5,9 @@ Authors: Andrew Yang, David Swinarski -/ import Mathlib.Algebra.Module.FinitePresentation import Mathlib.Algebra.Module.Projective -import Mathlib.LinearAlgebra.FreeModule.Basic +import Mathlib.LinearAlgebra.Dimension.Constructions +import Mathlib.LinearAlgebra.FreeModule.StrongRankCondition import Mathlib.RingTheory.LocalProperties.Submodule -import Mathlib.RingTheory.Localization.BaseChange /-! @@ -25,8 +25,43 @@ import Mathlib.RingTheory.Localization.BaseChange -/ -variable {R M N N'} [CommRing R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] -variable [AddCommGroup N'] [Module R N'] (S : Submonoid R) +universe uM + +variable {R N N' : Type*} {M : Type uM} [CommRing R] [AddCommGroup M] [Module R M] [AddCommGroup N] +variable [Module R N] [AddCommGroup N'] [Module R N'] (S : Submonoid R) + +theorem Module.free_of_isLocalizedModule {Rₛ Mₛ} [AddCommGroup Mₛ] [Module R Mₛ] + [CommRing Rₛ] [Algebra R Rₛ] [Module Rₛ Mₛ] [IsScalarTower R Rₛ Mₛ] + (S) (f : M →ₗ[R] Mₛ) [IsLocalization S Rₛ] [IsLocalizedModule S f] [Module.Free R M] : + Module.Free Rₛ Mₛ := + Free.of_equiv (IsLocalizedModule.isBaseChange S Rₛ f).equiv + +universe uR' uM' in +/-- +Also see `IsLocalizedModule.lift_rank_eq` for a version for non-free modules, +but requires `S` to not contain any zero-divisors. +-/ +theorem Module.lift_rank_of_isLocalizedModule_of_free + (Rₛ : Type uR') {Mₛ : Type uM'} [AddCommGroup Mₛ] [Module R Mₛ] + [CommRing Rₛ] [Algebra R Rₛ] [Module Rₛ Mₛ] [IsScalarTower R Rₛ Mₛ] (S : Submonoid R) + (f : M →ₗ[R] Mₛ) [IsLocalization S Rₛ] [IsLocalizedModule S f] [Module.Free R M] + [Nontrivial Rₛ] : + Cardinal.lift.{uM} (Module.rank Rₛ Mₛ) = Cardinal.lift.{uM'} (Module.rank R M) := by + apply Cardinal.lift_injective.{max uM' uR'} + have := (algebraMap R Rₛ).domain_nontrivial + have := (IsLocalizedModule.isBaseChange S Rₛ f).equiv.lift_rank_eq.symm + simp only [rank_tensorProduct, rank_self, + Cardinal.lift_one, one_mul, Cardinal.lift_lift] at this ⊢ + convert this + exact Cardinal.lift_umax + +theorem Module.finrank_of_isLocalizedModule_of_free + (Rₛ : Type*) {Mₛ : Type*} [AddCommGroup Mₛ] [Module R Mₛ] + [CommRing Rₛ] [Algebra R Rₛ] [Module Rₛ Mₛ] [IsScalarTower R Rₛ Mₛ] (S : Submonoid R) + (f : M →ₗ[R] Mₛ) [IsLocalization S Rₛ] [IsLocalizedModule S f] [Module.Free R M] + [Nontrivial Rₛ] : + Module.finrank Rₛ Mₛ = Module.finrank R M := by + simpa using congr(Cardinal.toNat $(Module.lift_rank_of_isLocalizedModule_of_free Rₛ S f)) theorem Module.projective_of_isLocalizedModule {Rₛ Mₛ} [AddCommGroup Mₛ] [Module R Mₛ] [CommRing Rₛ] [Algebra R Rₛ] [Module Rₛ Mₛ] [IsScalarTower R Rₛ Mₛ] From 454bc8a132d5ad572487513d801e41fc162dc128 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 25 Nov 2024 09:04:44 +0000 Subject: [PATCH 268/829] chore: split Order.Filter.Basic, creating Order.Filter.Finite (#19354) --- Mathlib.lean | 1 + Mathlib/Order/Filter/Bases.lean | 2 +- Mathlib/Order/Filter/Basic.lean | 311 +-------------- Mathlib/Order/Filter/CardinalInter.lean | 1 + Mathlib/Order/Filter/Cocardinal.lean | 1 - .../Order/Filter/CountableSeparatingOn.lean | 2 +- Mathlib/Order/Filter/Extr.lean | 1 + Mathlib/Order/Filter/Finite.lean | 359 ++++++++++++++++++ Mathlib/Order/Filter/Germ/Basic.lean | 1 + Mathlib/Order/Filter/Subsingleton.lean | 1 - .../SetTheory/Cardinal/CountableCover.lean | 2 +- 11 files changed, 380 insertions(+), 302 deletions(-) create mode 100644 Mathlib/Order/Filter/Finite.lean diff --git a/Mathlib.lean b/Mathlib.lean index 7e32e8e68409c..56e4a03916d8e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3953,6 +3953,7 @@ import Mathlib.Order.Filter.ENNReal import Mathlib.Order.Filter.EventuallyConst import Mathlib.Order.Filter.Extr import Mathlib.Order.Filter.FilterProduct +import Mathlib.Order.Filter.Finite import Mathlib.Order.Filter.Germ.Basic import Mathlib.Order.Filter.Germ.OrderedMonoid import Mathlib.Order.Filter.IndicatorFunction diff --git a/Mathlib/Order/Filter/Bases.lean b/Mathlib/Order/Filter/Bases.lean index ffd5fe2a6c9e4..ab8dd13872eb6 100644 --- a/Mathlib/Order/Filter/Bases.lean +++ b/Mathlib/Order/Filter/Bases.lean @@ -5,7 +5,7 @@ Authors: Yury Kudryashov, Johannes Hölzl, Mario Carneiro, Patrick Massot -/ import Mathlib.Data.Prod.PProd import Mathlib.Data.Set.Countable -import Mathlib.Order.Filter.Basic +import Mathlib.Order.Filter.Finite /-! # Filter bases diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index 995ddfea3234b..65c9595abbf40 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -3,7 +3,10 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Jeremy Avigad -/ -import Mathlib.Data.Set.Finite.Lattice +import Mathlib.Algebra.Group.Basic +import Mathlib.Algebra.Group.Pi.Basic +import Mathlib.Control.Basic +import Mathlib.Data.Set.Lattice import Mathlib.Order.Filter.Defs /-! @@ -59,6 +62,7 @@ we do *not* require. This gives `Filter X` better formal properties, in particul -/ assert_not_exists OrderedSemiring +assert_not_exists Fintype open Function Set Order open scoped symmDiff @@ -103,27 +107,15 @@ theorem congr_sets (h : { x | x ∈ s ↔ x ∈ t } ∈ f) : s ∈ f ↔ t ∈ f lemma copy_eq {S} (hmem : ∀ s, s ∈ S ↔ s ∈ f) : f.copy S hmem = f := Filter.ext hmem -@[simp] -theorem biInter_mem {β : Type v} {s : β → Set α} {is : Set β} (hf : is.Finite) : - (⋂ i ∈ is, s i) ∈ f ↔ ∀ i ∈ is, s i ∈ f := - Finite.induction_on hf (by simp) fun _ _ hs => by simp [hs] - -@[simp] -theorem biInter_finset_mem {β : Type v} {s : β → Set α} (is : Finset β) : - (⋂ i ∈ is, s i) ∈ f ↔ ∀ i ∈ is, s i ∈ f := - biInter_mem is.finite_toSet - -alias _root_.Finset.iInter_mem_sets := biInter_finset_mem +/-- Weaker version of `Filter.biInter_mem` that assumes `Subsingleton β` rather than `Finite β`. -/ +theorem biInter_mem' {β : Type v} {s : β → Set α} {is : Set β} (hf : is.Subsingleton) : + (⋂ i ∈ is, s i) ∈ f ↔ ∀ i ∈ is, s i ∈ f := by + apply Subsingleton.induction_on hf <;> simp --- attribute [protected] Finset.iInter_mem_sets porting note: doesn't work - -@[simp] -theorem sInter_mem {s : Set (Set α)} (hfin : s.Finite) : ⋂₀ s ∈ f ↔ ∀ U ∈ s, U ∈ f := by - rw [sInter_eq_biInter, biInter_mem hfin] - -@[simp] -theorem iInter_mem {β : Sort v} {s : β → Set α} [Finite β] : (⋂ i, s i) ∈ f ↔ ∀ i, s i ∈ f := - (sInter_mem (finite_range _)).trans forall_mem_range +/-- Weaker version of `Filter.iInter_mem` that assumes `Subsingleton β` rather than `Finite β`. -/ +theorem iInter_mem' {β : Sort v} {s : β → Set α} [Subsingleton β] : + (⋂ i, s i) ∈ f ↔ ∀ i, s i ∈ f := by + rw [← sInter_range, sInter_eq_biInter, biInter_mem' (subsingleton_range s), forall_mem_range] theorem exists_mem_subset_iff : (∃ t ∈ f, t ⊆ s) ↔ s ∈ f := ⟨fun ⟨_, ht, ts⟩ => mem_of_superset ht ts, fun hs => ⟨s, hs, Subset.rfl⟩⟩ @@ -181,24 +173,6 @@ theorem le_generate_iff {s : Set (Set α)} {f : Filter α} : f ≤ generate s hu.recOn (fun h' => h h') univ_mem (fun _ hxy hx => mem_of_superset hx hxy) fun _ _ hx hy => inter_mem hx hy -theorem mem_generate_iff {s : Set <| Set α} {U : Set α} : - U ∈ generate s ↔ ∃ t ⊆ s, Set.Finite t ∧ ⋂₀ t ⊆ U := by - constructor <;> intro h - · induction h with - | @basic V V_in => - exact ⟨{V}, singleton_subset_iff.2 V_in, finite_singleton _, (sInter_singleton _).subset⟩ - | univ => exact ⟨∅, empty_subset _, finite_empty, subset_univ _⟩ - | superset _ hVW hV => - rcases hV with ⟨t, hts, ht, htV⟩ - exact ⟨t, hts, ht, htV.trans hVW⟩ - | inter _ _ hV hW => - rcases hV, hW with ⟨⟨t, hts, ht, htV⟩, u, hus, hu, huW⟩ - exact - ⟨t ∪ u, union_subset hts hus, ht.union hu, - (sInter_union _ _).subset.trans <| inter_subset_inter htV huW⟩ - · rcases h with ⟨t, hts, tfin, h⟩ - exact mem_of_superset ((sInter_mem tfin).2 fun V hV => GenerateSets.basic <| hts hV) h - @[simp] lemma generate_singleton (s : Set α) : generate {s} = 𝓟 s := le_antisymm (fun _t ht ↦ mem_of_superset (mem_generate_of_mem <| mem_singleton _) ht) <| le_generate_iff.2 <| singleton_subset_iff.2 Subset.rfl @@ -334,58 +308,6 @@ theorem iInf_eq_generate (s : ι → Filter α) : iInf s = generate (⋃ i, (s i theorem mem_iInf_of_mem {f : ι → Filter α} (i : ι) {s} (hs : s ∈ f i) : s ∈ ⨅ i, f i := iInf_le f i hs -theorem mem_iInf_of_iInter {ι} {s : ι → Filter α} {U : Set α} {I : Set ι} (I_fin : I.Finite) - {V : I → Set α} (hV : ∀ (i : I), V i ∈ s i) (hU : ⋂ i, V i ⊆ U) : U ∈ ⨅ i, s i := by - haveI := I_fin.fintype - refine mem_of_superset (iInter_mem.2 fun i => ?_) hU - exact mem_iInf_of_mem (i : ι) (hV _) - -theorem mem_iInf {ι} {s : ι → Filter α} {U : Set α} : - (U ∈ ⨅ i, s i) ↔ - ∃ I : Set ι, I.Finite ∧ ∃ V : I → Set α, (∀ (i : I), V i ∈ s i) ∧ U = ⋂ i, V i := by - constructor - · rw [iInf_eq_generate, mem_generate_iff] - rintro ⟨t, tsub, tfin, tinter⟩ - rcases eq_finite_iUnion_of_finite_subset_iUnion tfin tsub with ⟨I, Ifin, σ, σfin, σsub, rfl⟩ - rw [sInter_iUnion] at tinter - set V := fun i => U ∪ ⋂₀ σ i with hV - have V_in : ∀ (i : I), V i ∈ s i := by - rintro i - have : ⋂₀ σ i ∈ s i := by - rw [sInter_mem (σfin _)] - apply σsub - exact mem_of_superset this subset_union_right - refine ⟨I, Ifin, V, V_in, ?_⟩ - rwa [hV, ← union_iInter, union_eq_self_of_subset_right] - · rintro ⟨I, Ifin, V, V_in, rfl⟩ - exact mem_iInf_of_iInter Ifin V_in Subset.rfl - -theorem mem_iInf' {ι} {s : ι → Filter α} {U : Set α} : - (U ∈ ⨅ i, s i) ↔ - ∃ I : Set ι, I.Finite ∧ ∃ V : ι → Set α, (∀ i, V i ∈ s i) ∧ - (∀ i ∉ I, V i = univ) ∧ (U = ⋂ i ∈ I, V i) ∧ U = ⋂ i, V i := by - classical - simp only [mem_iInf, SetCoe.forall', biInter_eq_iInter] - refine ⟨?_, fun ⟨I, If, V, hVs, _, hVU, _⟩ => ⟨I, If, fun i => V i, fun i => hVs i, hVU⟩⟩ - rintro ⟨I, If, V, hV, rfl⟩ - refine ⟨I, If, fun i => if hi : i ∈ I then V ⟨i, hi⟩ else univ, fun i => ?_, fun i hi => ?_, ?_⟩ - · dsimp only - split_ifs - exacts [hV ⟨i,_⟩, univ_mem] - · exact dif_neg hi - · simp only [iInter_dite, biInter_eq_iInter, dif_pos (Subtype.coe_prop _), Subtype.coe_eta, - iInter_univ, inter_univ, eq_self_iff_true, true_and] - -theorem exists_iInter_of_mem_iInf {ι : Type*} {α : Type*} {f : ι → Filter α} {s} - (hs : s ∈ ⨅ i, f i) : ∃ t : ι → Set α, (∀ i, t i ∈ f i) ∧ s = ⋂ i, t i := - let ⟨_, _, V, hVs, _, _, hVU'⟩ := mem_iInf'.1 hs; ⟨V, hVs, hVU'⟩ - -theorem mem_iInf_of_finite {ι : Type*} [Finite ι] {α : Type*} {f : ι → Filter α} (s) : - (s ∈ ⨅ i, f i) ↔ ∃ t : ι → Set α, (∀ i, t i ∈ f i) ∧ s = ⋂ i, t i := by - refine ⟨exists_iInter_of_mem_iInf, ?_⟩ - rintro ⟨t, ht, rfl⟩ - exact iInter_mem.2 fun i => mem_iInf_of_mem i (ht i) - @[simp] theorem le_principal_iff {s : Set α} {f : Filter α} : f ≤ 𝓟 s ↔ s ∈ f := ⟨fun h => h Subset.rfl, fun hs _ ht => mem_of_superset hs ht⟩ @@ -453,26 +375,6 @@ theorem NeBot.not_disjoint (hf : f.NeBot) (hs : s ∈ f) (ht : t ∈ f) : ¬Disj theorem inf_eq_bot_iff {f g : Filter α} : f ⊓ g = ⊥ ↔ ∃ U ∈ f, ∃ V ∈ g, U ∩ V = ∅ := by simp only [← disjoint_iff, Filter.disjoint_iff, Set.disjoint_iff_inter_eq_empty] -theorem _root_.Pairwise.exists_mem_filter_of_disjoint {ι : Type*} [Finite ι] {l : ι → Filter α} - (hd : Pairwise (Disjoint on l)) : - ∃ s : ι → Set α, (∀ i, s i ∈ l i) ∧ Pairwise (Disjoint on s) := by - have : Pairwise fun i j => ∃ (s : {s // s ∈ l i}) (t : {t // t ∈ l j}), Disjoint s.1 t.1 := by - simpa only [Pairwise, Function.onFun, Filter.disjoint_iff, exists_prop, Subtype.exists] using hd - choose! s t hst using this - refine ⟨fun i => ⋂ j, @s i j ∩ @t j i, fun i => ?_, fun i j hij => ?_⟩ - exacts [iInter_mem.2 fun j => inter_mem (@s i j).2 (@t j i).2, - (hst hij).mono ((iInter_subset _ j).trans inter_subset_left) - ((iInter_subset _ i).trans inter_subset_right)] - -theorem _root_.Set.PairwiseDisjoint.exists_mem_filter {ι : Type*} {l : ι → Filter α} {t : Set ι} - (hd : t.PairwiseDisjoint l) (ht : t.Finite) : - ∃ s : ι → Set α, (∀ i, s i ∈ l i) ∧ t.PairwiseDisjoint s := by - haveI := ht.to_subtype - rcases (hd.subtype _ _).exists_mem_filter_of_disjoint with ⟨s, hsl, hsd⟩ - lift s to (i : t) → {s // s ∈ l i} using hsl - rcases @Subtype.exists_pi_extension ι (fun i => { s // s ∈ l i }) _ _ s with ⟨s, rfl⟩ - exact ⟨fun i => s i, fun i => (s i).2, hsd.set_of_subtype _ _⟩ - /-- There is exactly one filter on an empty type. -/ instance unique [IsEmpty α] : Unique (Filter α) where default := ⊥ @@ -548,23 +450,6 @@ theorem biInf_sets_eq {f : β → Filter α} {s : Set β} (h : DirectedOn (f ⁻ (ne : s.Nonempty) : (⨅ i ∈ s, f i).sets = ⋃ i ∈ s, (f i).sets := ext fun t => by simp [mem_biInf_of_directed h ne] -theorem iInf_sets_eq_finite {ι : Type*} (f : ι → Filter α) : - (⨅ i, f i).sets = ⋃ t : Finset ι, (⨅ i ∈ t, f i).sets := by - rw [iInf_eq_iInf_finset, iInf_sets_eq] - exact directed_of_isDirected_le fun _ _ => biInf_mono - -theorem iInf_sets_eq_finite' (f : ι → Filter α) : - (⨅ i, f i).sets = ⋃ t : Finset (PLift ι), (⨅ i ∈ t, f (PLift.down i)).sets := by - rw [← iInf_sets_eq_finite, ← Equiv.plift.surjective.iInf_comp, Equiv.plift_apply] - -theorem mem_iInf_finite {ι : Type*} {f : ι → Filter α} (s) : - s ∈ iInf f ↔ ∃ t : Finset ι, s ∈ ⨅ i ∈ t, f i := - (Set.ext_iff.1 (iInf_sets_eq_finite f) s).trans mem_iUnion - -theorem mem_iInf_finite' {f : ι → Filter α} (s) : - s ∈ iInf f ↔ ∃ t : Finset (PLift ι), s ∈ ⨅ i ∈ t, f (PLift.down i) := - (Set.ext_iff.1 (iInf_sets_eq_finite' f) s).trans mem_iUnion - @[simp] theorem sup_join {f₁ f₂ : Filter (Filter α)} : join f₁ ⊔ join f₂ = join (f₁ ⊔ f₂) := Filter.ext fun x => by simp only [mem_sup, mem_join] @@ -583,38 +468,6 @@ instance : DistribLattice (Filter α) := ⟨t₁, x.sets_of_superset hs inter_subset_left, ht₁, t₂, x.sets_of_superset hs inter_subset_right, ht₂, rfl⟩ } -/-- The dual version does not hold! `Filter α` is not a `CompleteDistribLattice`. -/ --- See note [reducible non-instances] -abbrev coframeMinimalAxioms : Coframe.MinimalAxioms (Filter α) := - { Filter.instCompleteLatticeFilter with - iInf_sup_le_sup_sInf := fun f s t ⟨h₁, h₂⟩ => by - classical - rw [iInf_subtype'] - rw [sInf_eq_iInf', ← Filter.mem_sets, iInf_sets_eq_finite, mem_iUnion] at h₂ - obtain ⟨u, hu⟩ := h₂ - rw [← Finset.inf_eq_iInf] at hu - suffices ⨅ i : s, f ⊔ ↑i ≤ f ⊔ u.inf fun i => ↑i from this ⟨h₁, hu⟩ - refine Finset.induction_on u (le_sup_of_le_right le_top) ?_ - rintro ⟨i⟩ u _ ih - rw [Finset.inf_insert, sup_inf_left] - exact le_inf (iInf_le _ _) ih } - -instance instCoframe : Coframe (Filter α) := .ofMinimalAxioms coframeMinimalAxioms - -theorem mem_iInf_finset {s : Finset α} {f : α → Filter β} {t : Set β} : - (t ∈ ⨅ a ∈ s, f a) ↔ ∃ p : α → Set β, (∀ a ∈ s, p a ∈ f a) ∧ t = ⋂ a ∈ s, p a := by - classical - simp only [← Finset.set_biInter_coe, biInter_eq_iInter, iInf_subtype'] - refine ⟨fun h => ?_, ?_⟩ - · rcases (mem_iInf_of_finite _).1 h with ⟨p, hp, rfl⟩ - refine ⟨fun a => if h : a ∈ s then p ⟨a, h⟩ else univ, - fun a ha => by simpa [ha] using hp ⟨a, ha⟩, ?_⟩ - refine iInter_congr_of_surjective id surjective_id ?_ - rintro ⟨a, ha⟩ - simp [ha] - · rintro ⟨p, hpf, rfl⟩ - exact iInter_mem.2 fun a => mem_iInf_of_mem a (hpf a a.2) - /-- If `f : ι → Filter α` is directed, `ι` is not empty, and `∀ i, f i ≠ ⊥`, then `iInf f ≠ ⊥`. See also `iInf_neBot_of_directed` for a version assuming `Nonempty α` instead of `Nonempty ι`. -/ theorem iInf_neBot_of_directed' {f : ι → Filter α} [Nonempty ι] (hd : Directed (· ≥ ·) f) : @@ -650,20 +503,6 @@ theorem iInf_neBot_iff_of_directed {f : ι → Filter α} [Nonempty α] (hd : Di NeBot (iInf f) ↔ ∀ i, NeBot (f i) := ⟨fun H i => H.mono (iInf_le _ i), iInf_neBot_of_directed hd⟩ -@[elab_as_elim] -theorem iInf_sets_induct {f : ι → Filter α} {s : Set α} (hs : s ∈ iInf f) {p : Set α → Prop} - (uni : p univ) (ins : ∀ {i s₁ s₂}, s₁ ∈ f i → p s₂ → p (s₁ ∩ s₂)) : p s := by - classical - rw [mem_iInf_finite'] at hs - simp only [← Finset.inf_eq_iInf] at hs - rcases hs with ⟨is, his⟩ - induction is using Finset.induction_on generalizing s with - | empty => rwa [mem_top.1 his] - | insert _ ih => - rw [Finset.inf_insert, mem_inf_iff] at his - rcases his with ⟨s₁, hs₁, s₂, hs₂, rfl⟩ - exact ins hs₁ (ih hs₂) - /-! #### `principal` equations -/ @[simp] @@ -719,29 +558,6 @@ theorem diff_mem_inf_principal_compl {f : Filter α} {s : Set α} (hs : s ∈ f) theorem principal_le_iff {s : Set α} {f : Filter α} : 𝓟 s ≤ f ↔ ∀ V ∈ f, s ⊆ V := by simp_rw [le_def, mem_principal] -@[simp] -theorem iInf_principal_finset {ι : Type w} (s : Finset ι) (f : ι → Set α) : - ⨅ i ∈ s, 𝓟 (f i) = 𝓟 (⋂ i ∈ s, f i) := by - classical - induction' s using Finset.induction_on with i s _ hs - · simp - · rw [Finset.iInf_insert, Finset.set_biInter_insert, hs, inf_principal] - -theorem iInf_principal {ι : Sort w} [Finite ι] (f : ι → Set α) : ⨅ i, 𝓟 (f i) = 𝓟 (⋂ i, f i) := by - cases nonempty_fintype (PLift ι) - rw [← iInf_plift_down, ← iInter_plift_down] - simpa using iInf_principal_finset Finset.univ (f <| PLift.down ·) - -/-- A special case of `iInf_principal` that is safe to mark `simp`. -/ -@[simp] -theorem iInf_principal' {ι : Type w} [Finite ι] (f : ι → Set α) : ⨅ i, 𝓟 (f i) = 𝓟 (⋂ i, f i) := - iInf_principal _ - -theorem iInf_principal_finite {ι : Type w} {s : Set ι} (hs : s.Finite) (f : ι → Set α) : - ⨅ i ∈ s, 𝓟 (f i) = 𝓟 (⋂ i ∈ s, f i) := by - lift s to Finset ι using hs - exact mod_cast iInf_principal_finset s f - end Lattice @[mono, gcongr] @@ -820,28 +636,6 @@ theorem eventually_congr {f : Filter α} {p q : α → Prop} (h : ∀ᶠ x in f, (∀ᶠ x in f, p x) ↔ ∀ᶠ x in f, q x := ⟨fun hp => hp.congr h, fun hq => hq.congr <| by simpa only [Iff.comm] using h⟩ -@[simp] -theorem eventually_all {ι : Sort*} [Finite ι] {l} {p : ι → α → Prop} : - (∀ᶠ x in l, ∀ i, p i x) ↔ ∀ i, ∀ᶠ x in l, p i x := by - simpa only [Filter.Eventually, setOf_forall] using iInter_mem - -@[simp] -theorem eventually_all_finite {ι} {I : Set ι} (hI : I.Finite) {l} {p : ι → α → Prop} : - (∀ᶠ x in l, ∀ i ∈ I, p i x) ↔ ∀ i ∈ I, ∀ᶠ x in l, p i x := by - simpa only [Filter.Eventually, setOf_forall] using biInter_mem hI - -alias _root_.Set.Finite.eventually_all := eventually_all_finite - --- attribute [protected] Set.Finite.eventually_all - -@[simp] theorem eventually_all_finset {ι} (I : Finset ι) {l} {p : ι → α → Prop} : - (∀ᶠ x in l, ∀ i ∈ I, p i x) ↔ ∀ i ∈ I, ∀ᶠ x in l, p i x := - I.finite_toSet.eventually_all - -alias _root_.Finset.eventually_all := eventually_all_finset - --- attribute [protected] Finset.eventually_all - @[simp] theorem eventually_or_distrib_left {f : Filter α} {p : Prop} {q : α → Prop} : (∀ᶠ x in f, p ∨ q x) ↔ p ∨ ∀ᶠ x in f, q x := @@ -852,10 +646,6 @@ theorem eventually_or_distrib_right {f : Filter α} {p : α → Prop} {q : Prop} (∀ᶠ x in f, p x ∨ q) ↔ (∀ᶠ x in f, p x) ∨ q := by simp only [@or_comm _ q, eventually_or_distrib_left] -theorem eventually_imp_distrib_left {f : Filter α} {p : Prop} {q : α → Prop} : - (∀ᶠ x in f, p → q x) ↔ p → ∀ᶠ x in f, q x := - eventually_all - @[simp] theorem eventually_bot {p : α → Prop} : ∀ᶠ x in ⊥, p x := ⟨⟩ @@ -1008,16 +798,6 @@ theorem eventually_imp_distrib_right {f : Filter α} {p : α → Prop} {q : Prop (∀ᶠ x in f, p x → q) ↔ (∃ᶠ x in f, p x) → q := by simp only [imp_iff_not_or, eventually_or_distrib_right, not_frequently] -@[simp] -theorem frequently_and_distrib_left {f : Filter α} {p : Prop} {q : α → Prop} : - (∃ᶠ x in f, p ∧ q x) ↔ p ∧ ∃ᶠ x in f, q x := by - simp only [Filter.Frequently, not_and, eventually_imp_distrib_left, Classical.not_imp] - -@[simp] -theorem frequently_and_distrib_right {f : Filter α} {p : α → Prop} {q : Prop} : - (∃ᶠ x in f, p x ∧ q) ↔ (∃ᶠ x in f, p x) ∧ q := by - simp only [@and_comm _ q, frequently_and_distrib_left] - @[simp] theorem frequently_bot {p : α → Prop} : ¬∃ᶠ x in ⊥, p x := by simp @@ -1328,69 +1108,6 @@ theorem EventuallyLE.union {s t s' t' : Set α} {l : Filter α} (h : s ≤ᶠ[l] (s ∪ s' : Set α) ≤ᶠ[l] (t ∪ t' : Set α) := h'.mp <| h.mono fun _ => Or.imp -protected lemma EventuallyLE.iUnion [Finite ι] {s t : ι → Set α} - (h : ∀ i, s i ≤ᶠ[l] t i) : (⋃ i, s i) ≤ᶠ[l] ⋃ i, t i := - (eventually_all.2 h).mono fun _x hx hx' ↦ - let ⟨i, hi⟩ := mem_iUnion.1 hx'; mem_iUnion.2 ⟨i, hx i hi⟩ - -protected lemma EventuallyEq.iUnion [Finite ι] {s t : ι → Set α} - (h : ∀ i, s i =ᶠ[l] t i) : (⋃ i, s i) =ᶠ[l] ⋃ i, t i := - (EventuallyLE.iUnion fun i ↦ (h i).le).antisymm <| .iUnion fun i ↦ (h i).symm.le - -protected lemma EventuallyLE.iInter [Finite ι] {s t : ι → Set α} - (h : ∀ i, s i ≤ᶠ[l] t i) : (⋂ i, s i) ≤ᶠ[l] ⋂ i, t i := - (eventually_all.2 h).mono fun _x hx hx' ↦ mem_iInter.2 fun i ↦ hx i (mem_iInter.1 hx' i) - -protected lemma EventuallyEq.iInter [Finite ι] {s t : ι → Set α} - (h : ∀ i, s i =ᶠ[l] t i) : (⋂ i, s i) =ᶠ[l] ⋂ i, t i := - (EventuallyLE.iInter fun i ↦ (h i).le).antisymm <| .iInter fun i ↦ (h i).symm.le - -lemma _root_.Set.Finite.eventuallyLE_iUnion {ι : Type*} {s : Set ι} (hs : s.Finite) - {f g : ι → Set α} (hle : ∀ i ∈ s, f i ≤ᶠ[l] g i) : (⋃ i ∈ s, f i) ≤ᶠ[l] (⋃ i ∈ s, g i) := by - have := hs.to_subtype - rw [biUnion_eq_iUnion, biUnion_eq_iUnion] - exact .iUnion fun i ↦ hle i.1 i.2 - -alias EventuallyLE.biUnion := Set.Finite.eventuallyLE_iUnion - -lemma _root_.Set.Finite.eventuallyEq_iUnion {ι : Type*} {s : Set ι} (hs : s.Finite) - {f g : ι → Set α} (heq : ∀ i ∈ s, f i =ᶠ[l] g i) : (⋃ i ∈ s, f i) =ᶠ[l] (⋃ i ∈ s, g i) := - (EventuallyLE.biUnion hs fun i hi ↦ (heq i hi).le).antisymm <| - .biUnion hs fun i hi ↦ (heq i hi).symm.le - -alias EventuallyEq.biUnion := Set.Finite.eventuallyEq_iUnion - -lemma _root_.Set.Finite.eventuallyLE_iInter {ι : Type*} {s : Set ι} (hs : s.Finite) - {f g : ι → Set α} (hle : ∀ i ∈ s, f i ≤ᶠ[l] g i) : (⋂ i ∈ s, f i) ≤ᶠ[l] (⋂ i ∈ s, g i) := by - have := hs.to_subtype - rw [biInter_eq_iInter, biInter_eq_iInter] - exact .iInter fun i ↦ hle i.1 i.2 - -alias EventuallyLE.biInter := Set.Finite.eventuallyLE_iInter - -lemma _root_.Set.Finite.eventuallyEq_iInter {ι : Type*} {s : Set ι} (hs : s.Finite) - {f g : ι → Set α} (heq : ∀ i ∈ s, f i =ᶠ[l] g i) : (⋂ i ∈ s, f i) =ᶠ[l] (⋂ i ∈ s, g i) := - (EventuallyLE.biInter hs fun i hi ↦ (heq i hi).le).antisymm <| - .biInter hs fun i hi ↦ (heq i hi).symm.le - -alias EventuallyEq.biInter := Set.Finite.eventuallyEq_iInter - -lemma _root_.Finset.eventuallyLE_iUnion {ι : Type*} (s : Finset ι) {f g : ι → Set α} - (hle : ∀ i ∈ s, f i ≤ᶠ[l] g i) : (⋃ i ∈ s, f i) ≤ᶠ[l] (⋃ i ∈ s, g i) := - .biUnion s.finite_toSet hle - -lemma _root_.Finset.eventuallyEq_iUnion {ι : Type*} (s : Finset ι) {f g : ι → Set α} - (heq : ∀ i ∈ s, f i =ᶠ[l] g i) : (⋃ i ∈ s, f i) =ᶠ[l] (⋃ i ∈ s, g i) := - .biUnion s.finite_toSet heq - -lemma _root_.Finset.eventuallyLE_iInter {ι : Type*} (s : Finset ι) {f g : ι → Set α} - (hle : ∀ i ∈ s, f i ≤ᶠ[l] g i) : (⋂ i ∈ s, f i) ≤ᶠ[l] (⋂ i ∈ s, g i) := - .biInter s.finite_toSet hle - -lemma _root_.Finset.eventuallyEq_iInter {ι : Type*} (s : Finset ι) {f g : ι → Set α} - (heq : ∀ i ∈ s, f i =ᶠ[l] g i) : (⋂ i ∈ s, f i) =ᶠ[l] (⋂ i ∈ s, g i) := - .biInter s.finite_toSet heq - @[mono] theorem EventuallyLE.compl {s t : Set α} {l : Filter α} (h : s ≤ᶠ[l] t) : (tᶜ : Set α) ≤ᶠ[l] (sᶜ : Set α) := @@ -2372,4 +2089,4 @@ lemma compl_mem_comk {p : Set α → Prop} {he hmono hunion s} : end Filter -set_option linter.style.longFile 2500 +set_option linter.style.longFile 2200 diff --git a/Mathlib/Order/Filter/CardinalInter.lean b/Mathlib/Order/Filter/CardinalInter.lean index 22967d8a816b3..9a1e5bb744169 100644 --- a/Mathlib/Order/Filter/CardinalInter.lean +++ b/Mathlib/Order/Filter/CardinalInter.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Josha Dekker -/ import Mathlib.Order.Filter.Tendsto +import Mathlib.Order.Filter.Finite import Mathlib.Order.Filter.CountableInter import Mathlib.SetTheory.Cardinal.Arithmetic import Mathlib.SetTheory.Cardinal.Cofinality diff --git a/Mathlib/Order/Filter/Cocardinal.lean b/Mathlib/Order/Filter/Cocardinal.lean index 25e10e20a8fa3..d3da9a3ef6505 100644 --- a/Mathlib/Order/Filter/Cocardinal.lean +++ b/Mathlib/Order/Filter/Cocardinal.lean @@ -8,7 +8,6 @@ import Mathlib.Order.Filter.CountableInter import Mathlib.Order.Filter.CardinalInter import Mathlib.SetTheory.Cardinal.Arithmetic import Mathlib.SetTheory.Cardinal.Cofinality -import Mathlib.Order.Filter.Bases /-! # The cocardinal filter diff --git a/Mathlib/Order/Filter/CountableSeparatingOn.lean b/Mathlib/Order/Filter/CountableSeparatingOn.lean index b1f268cda8419..552a128f17bc5 100644 --- a/Mathlib/Order/Filter/CountableSeparatingOn.lean +++ b/Mathlib/Order/Filter/CountableSeparatingOn.lean @@ -167,7 +167,7 @@ theorem exists_subset_subsingleton_mem_of_forall_separating (p : Set α → Prop | inr hsl => simp only [hx.2 s hsS hsl, hy.2 s hsS hsl] · exact inter_mem (inter_mem hs ((countable_sInter_mem (hSc.mono inter_subset_left)).2 fun _ h ↦ h.2)) - ((countable_bInter_mem hSc).2 fun U hU ↦ iInter_mem.2 id) + ((countable_bInter_mem hSc).2 fun U hU ↦ iInter_mem'.2 id) theorem exists_mem_singleton_mem_of_mem_of_nonempty_of_forall_separating (p : Set α → Prop) {s : Set α} [HasCountableSeparatingOn α p s] (hs : s ∈ l) (hne : s.Nonempty) diff --git a/Mathlib/Order/Filter/Extr.lean b/Mathlib/Order/Filter/Extr.lean index ba1e7c93d0896..5f405e77557b1 100644 --- a/Mathlib/Order/Filter/Extr.lean +++ b/Mathlib/Order/Filter/Extr.lean @@ -6,6 +6,7 @@ Authors: Yury Kudryashov import Mathlib.Order.Filter.Tendsto import Mathlib.Order.ConditionallyCompleteLattice.Indexed import Mathlib.Algebra.Order.Group.Defs +import Mathlib.Data.Finset.Lattice.Fold /-! # Minimum and maximum w.r.t. a filter and on a set diff --git a/Mathlib/Order/Filter/Finite.lean b/Mathlib/Order/Filter/Finite.lean new file mode 100644 index 0000000000000..2c96231b9a89f --- /dev/null +++ b/Mathlib/Order/Filter/Finite.lean @@ -0,0 +1,359 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Jeremy Avigad +-/ +import Mathlib.Data.Set.Finite.Lattice +import Mathlib.Order.Filter.Basic + +/-! +# Results filters related to finiteness. + +-/ + + + +open Function Set Order +open scoped symmDiff + +universe u v w x y + +namespace Filter + +variable {α : Type u} {f g : Filter α} {s t : Set α} + +@[simp] +theorem biInter_mem {β : Type v} {s : β → Set α} {is : Set β} (hf : is.Finite) : + (⋂ i ∈ is, s i) ∈ f ↔ ∀ i ∈ is, s i ∈ f := + Finite.induction_on hf (by simp) fun _ _ hs => by simp [hs] + +@[simp] +theorem biInter_finset_mem {β : Type v} {s : β → Set α} (is : Finset β) : + (⋂ i ∈ is, s i) ∈ f ↔ ∀ i ∈ is, s i ∈ f := + biInter_mem is.finite_toSet + +alias _root_.Finset.iInter_mem_sets := biInter_finset_mem + +-- attribute [protected] Finset.iInter_mem_sets porting note: doesn't work + +@[simp] +theorem sInter_mem {s : Set (Set α)} (hfin : s.Finite) : ⋂₀ s ∈ f ↔ ∀ U ∈ s, U ∈ f := by + rw [sInter_eq_biInter, biInter_mem hfin] + +@[simp] +theorem iInter_mem {β : Sort v} {s : β → Set α} [Finite β] : (⋂ i, s i) ∈ f ↔ ∀ i, s i ∈ f := + (sInter_mem (finite_range _)).trans forall_mem_range + +end Filter + + +namespace Filter + +variable {α : Type u} {β : Type v} {γ : Type w} {δ : Type*} {ι : Sort x} + +section Lattice + +variable {f g : Filter α} {s t : Set α} + +theorem mem_generate_iff {s : Set <| Set α} {U : Set α} : + U ∈ generate s ↔ ∃ t ⊆ s, Set.Finite t ∧ ⋂₀ t ⊆ U := by + constructor <;> intro h + · induction h with + | @basic V V_in => + exact ⟨{V}, singleton_subset_iff.2 V_in, finite_singleton _, (sInter_singleton _).subset⟩ + | univ => exact ⟨∅, empty_subset _, finite_empty, subset_univ _⟩ + | superset _ hVW hV => + rcases hV with ⟨t, hts, ht, htV⟩ + exact ⟨t, hts, ht, htV.trans hVW⟩ + | inter _ _ hV hW => + rcases hV, hW with ⟨⟨t, hts, ht, htV⟩, u, hus, hu, huW⟩ + exact + ⟨t ∪ u, union_subset hts hus, ht.union hu, + (sInter_union _ _).subset.trans <| inter_subset_inter htV huW⟩ + · rcases h with ⟨t, hts, tfin, h⟩ + exact mem_of_superset ((sInter_mem tfin).2 fun V hV => GenerateSets.basic <| hts hV) h + +theorem mem_iInf_of_iInter {ι} {s : ι → Filter α} {U : Set α} {I : Set ι} (I_fin : I.Finite) + {V : I → Set α} (hV : ∀ (i : I), V i ∈ s i) (hU : ⋂ i, V i ⊆ U) : U ∈ ⨅ i, s i := by + haveI := I_fin.fintype + refine mem_of_superset (iInter_mem.2 fun i => ?_) hU + exact mem_iInf_of_mem (i : ι) (hV _) + +theorem mem_iInf {ι} {s : ι → Filter α} {U : Set α} : + (U ∈ ⨅ i, s i) ↔ + ∃ I : Set ι, I.Finite ∧ ∃ V : I → Set α, (∀ (i : I), V i ∈ s i) ∧ U = ⋂ i, V i := by + constructor + · rw [iInf_eq_generate, mem_generate_iff] + rintro ⟨t, tsub, tfin, tinter⟩ + rcases eq_finite_iUnion_of_finite_subset_iUnion tfin tsub with ⟨I, Ifin, σ, σfin, σsub, rfl⟩ + rw [sInter_iUnion] at tinter + set V := fun i => U ∪ ⋂₀ σ i with hV + have V_in : ∀ (i : I), V i ∈ s i := by + rintro i + have : ⋂₀ σ i ∈ s i := by + rw [sInter_mem (σfin _)] + apply σsub + exact mem_of_superset this subset_union_right + refine ⟨I, Ifin, V, V_in, ?_⟩ + rwa [hV, ← union_iInter, union_eq_self_of_subset_right] + · rintro ⟨I, Ifin, V, V_in, rfl⟩ + exact mem_iInf_of_iInter Ifin V_in Subset.rfl + +theorem mem_iInf' {ι} {s : ι → Filter α} {U : Set α} : + (U ∈ ⨅ i, s i) ↔ + ∃ I : Set ι, I.Finite ∧ ∃ V : ι → Set α, (∀ i, V i ∈ s i) ∧ + (∀ i ∉ I, V i = univ) ∧ (U = ⋂ i ∈ I, V i) ∧ U = ⋂ i, V i := by + classical + simp only [mem_iInf, SetCoe.forall', biInter_eq_iInter] + refine ⟨?_, fun ⟨I, If, V, hVs, _, hVU, _⟩ => ⟨I, If, fun i => V i, fun i => hVs i, hVU⟩⟩ + rintro ⟨I, If, V, hV, rfl⟩ + refine ⟨I, If, fun i => if hi : i ∈ I then V ⟨i, hi⟩ else univ, fun i => ?_, fun i hi => ?_, ?_⟩ + · dsimp only + split_ifs + exacts [hV ⟨i,_⟩, univ_mem] + · exact dif_neg hi + · simp only [iInter_dite, biInter_eq_iInter, dif_pos (Subtype.coe_prop _), Subtype.coe_eta, + iInter_univ, inter_univ, eq_self_iff_true, true_and] + +theorem exists_iInter_of_mem_iInf {ι : Type*} {α : Type*} {f : ι → Filter α} {s} + (hs : s ∈ ⨅ i, f i) : ∃ t : ι → Set α, (∀ i, t i ∈ f i) ∧ s = ⋂ i, t i := + let ⟨_, _, V, hVs, _, _, hVU'⟩ := mem_iInf'.1 hs; ⟨V, hVs, hVU'⟩ + +theorem mem_iInf_of_finite {ι : Type*} [Finite ι] {α : Type*} {f : ι → Filter α} (s) : + (s ∈ ⨅ i, f i) ↔ ∃ t : ι → Set α, (∀ i, t i ∈ f i) ∧ s = ⋂ i, t i := by + refine ⟨exists_iInter_of_mem_iInf, ?_⟩ + rintro ⟨t, ht, rfl⟩ + exact iInter_mem.2 fun i => mem_iInf_of_mem i (ht i) + +/-! ### Lattice equations -/ + +theorem _root_.Pairwise.exists_mem_filter_of_disjoint {ι : Type*} [Finite ι] {l : ι → Filter α} + (hd : Pairwise (Disjoint on l)) : + ∃ s : ι → Set α, (∀ i, s i ∈ l i) ∧ Pairwise (Disjoint on s) := by + have : Pairwise fun i j => ∃ (s : {s // s ∈ l i}) (t : {t // t ∈ l j}), Disjoint s.1 t.1 := by + simpa only [Pairwise, Function.onFun, Filter.disjoint_iff, exists_prop, Subtype.exists] using hd + choose! s t hst using this + refine ⟨fun i => ⋂ j, @s i j ∩ @t j i, fun i => ?_, fun i j hij => ?_⟩ + exacts [iInter_mem.2 fun j => inter_mem (@s i j).2 (@t j i).2, + (hst hij).mono ((iInter_subset _ j).trans inter_subset_left) + ((iInter_subset _ i).trans inter_subset_right)] + +theorem _root_.Set.PairwiseDisjoint.exists_mem_filter {ι : Type*} {l : ι → Filter α} {t : Set ι} + (hd : t.PairwiseDisjoint l) (ht : t.Finite) : + ∃ s : ι → Set α, (∀ i, s i ∈ l i) ∧ t.PairwiseDisjoint s := by + haveI := ht.to_subtype + rcases (hd.subtype _ _).exists_mem_filter_of_disjoint with ⟨s, hsl, hsd⟩ + lift s to (i : t) → {s // s ∈ l i} using hsl + rcases @Subtype.exists_pi_extension ι (fun i => { s // s ∈ l i }) _ _ s with ⟨s, rfl⟩ + exact ⟨fun i => s i, fun i => (s i).2, hsd.set_of_subtype _ _⟩ + + +theorem iInf_sets_eq_finite {ι : Type*} (f : ι → Filter α) : + (⨅ i, f i).sets = ⋃ t : Finset ι, (⨅ i ∈ t, f i).sets := by + rw [iInf_eq_iInf_finset, iInf_sets_eq] + exact directed_of_isDirected_le fun _ _ => biInf_mono + +theorem iInf_sets_eq_finite' (f : ι → Filter α) : + (⨅ i, f i).sets = ⋃ t : Finset (PLift ι), (⨅ i ∈ t, f (PLift.down i)).sets := by + rw [← iInf_sets_eq_finite, ← Equiv.plift.surjective.iInf_comp, Equiv.plift_apply] + +theorem mem_iInf_finite {ι : Type*} {f : ι → Filter α} (s) : + s ∈ iInf f ↔ ∃ t : Finset ι, s ∈ ⨅ i ∈ t, f i := + (Set.ext_iff.1 (iInf_sets_eq_finite f) s).trans mem_iUnion + +theorem mem_iInf_finite' {f : ι → Filter α} (s) : + s ∈ iInf f ↔ ∃ t : Finset (PLift ι), s ∈ ⨅ i ∈ t, f (PLift.down i) := + (Set.ext_iff.1 (iInf_sets_eq_finite' f) s).trans mem_iUnion + +/-- The dual version does not hold! `Filter α` is not a `CompleteDistribLattice`. -/ +-- See note [reducible non-instances] +abbrev coframeMinimalAxioms : Coframe.MinimalAxioms (Filter α) := + { Filter.instCompleteLatticeFilter with + iInf_sup_le_sup_sInf := fun f s t ⟨h₁, h₂⟩ => by + classical + rw [iInf_subtype'] + rw [sInf_eq_iInf', ← Filter.mem_sets, iInf_sets_eq_finite, mem_iUnion] at h₂ + obtain ⟨u, hu⟩ := h₂ + rw [← Finset.inf_eq_iInf] at hu + suffices ⨅ i : s, f ⊔ ↑i ≤ f ⊔ u.inf fun i => ↑i from this ⟨h₁, hu⟩ + refine Finset.induction_on u (le_sup_of_le_right le_top) ?_ + rintro ⟨i⟩ u _ ih + rw [Finset.inf_insert, sup_inf_left] + exact le_inf (iInf_le _ _) ih } + +instance instCoframe : Coframe (Filter α) := .ofMinimalAxioms coframeMinimalAxioms + +theorem mem_iInf_finset {s : Finset α} {f : α → Filter β} {t : Set β} : + (t ∈ ⨅ a ∈ s, f a) ↔ ∃ p : α → Set β, (∀ a ∈ s, p a ∈ f a) ∧ t = ⋂ a ∈ s, p a := by + classical + simp only [← Finset.set_biInter_coe, biInter_eq_iInter, iInf_subtype'] + refine ⟨fun h => ?_, ?_⟩ + · rcases (mem_iInf_of_finite _).1 h with ⟨p, hp, rfl⟩ + refine ⟨fun a => if h : a ∈ s then p ⟨a, h⟩ else univ, + fun a ha => by simpa [ha] using hp ⟨a, ha⟩, ?_⟩ + refine iInter_congr_of_surjective id surjective_id ?_ + rintro ⟨a, ha⟩ + simp [ha] + · rintro ⟨p, hpf, rfl⟩ + exact iInter_mem.2 fun a => mem_iInf_of_mem a (hpf a a.2) + + +@[elab_as_elim] +theorem iInf_sets_induct {f : ι → Filter α} {s : Set α} (hs : s ∈ iInf f) {p : Set α → Prop} + (uni : p univ) (ins : ∀ {i s₁ s₂}, s₁ ∈ f i → p s₂ → p (s₁ ∩ s₂)) : p s := by + classical + rw [mem_iInf_finite'] at hs + simp only [← Finset.inf_eq_iInf] at hs + rcases hs with ⟨is, his⟩ + induction is using Finset.induction_on generalizing s with + | empty => rwa [mem_top.1 his] + | insert _ ih => + rw [Finset.inf_insert, mem_inf_iff] at his + rcases his with ⟨s₁, hs₁, s₂, hs₂, rfl⟩ + exact ins hs₁ (ih hs₂) + +/-! #### `principal` equations -/ + +@[simp] +theorem iInf_principal_finset {ι : Type w} (s : Finset ι) (f : ι → Set α) : + ⨅ i ∈ s, 𝓟 (f i) = 𝓟 (⋂ i ∈ s, f i) := by + classical + induction' s using Finset.induction_on with i s _ hs + · simp + · rw [Finset.iInf_insert, Finset.set_biInter_insert, hs, inf_principal] + +theorem iInf_principal {ι : Sort w} [Finite ι] (f : ι → Set α) : ⨅ i, 𝓟 (f i) = 𝓟 (⋂ i, f i) := by + cases nonempty_fintype (PLift ι) + rw [← iInf_plift_down, ← iInter_plift_down] + simpa using iInf_principal_finset Finset.univ (f <| PLift.down ·) + +/-- A special case of `iInf_principal` that is safe to mark `simp`. -/ +@[simp] +theorem iInf_principal' {ι : Type w} [Finite ι] (f : ι → Set α) : ⨅ i, 𝓟 (f i) = 𝓟 (⋂ i, f i) := + iInf_principal _ + +theorem iInf_principal_finite {ι : Type w} {s : Set ι} (hs : s.Finite) (f : ι → Set α) : + ⨅ i ∈ s, 𝓟 (f i) = 𝓟 (⋂ i ∈ s, f i) := by + lift s to Finset ι using hs + exact mod_cast iInf_principal_finset s f + +end Lattice + +/-! ### Eventually -/ + +@[simp] +theorem eventually_all {ι : Sort*} [Finite ι] {l} {p : ι → α → Prop} : + (∀ᶠ x in l, ∀ i, p i x) ↔ ∀ i, ∀ᶠ x in l, p i x := by + simpa only [Filter.Eventually, setOf_forall] using iInter_mem + +@[simp] +theorem eventually_all_finite {ι} {I : Set ι} (hI : I.Finite) {l} {p : ι → α → Prop} : + (∀ᶠ x in l, ∀ i ∈ I, p i x) ↔ ∀ i ∈ I, ∀ᶠ x in l, p i x := by + simpa only [Filter.Eventually, setOf_forall] using biInter_mem hI + +alias _root_.Set.Finite.eventually_all := eventually_all_finite + +-- attribute [protected] Set.Finite.eventually_all + +@[simp] theorem eventually_all_finset {ι} (I : Finset ι) {l} {p : ι → α → Prop} : + (∀ᶠ x in l, ∀ i ∈ I, p i x) ↔ ∀ i ∈ I, ∀ᶠ x in l, p i x := + I.finite_toSet.eventually_all + +alias _root_.Finset.eventually_all := eventually_all_finset + +-- attribute [protected] Finset.eventually_all + +theorem eventually_imp_distrib_left {f : Filter α} {p : Prop} {q : α → Prop} : + (∀ᶠ x in f, p → q x) ↔ p → ∀ᶠ x in f, q x := + eventually_all + + +/-! ### Frequently -/ + +@[simp] +theorem frequently_and_distrib_left {f : Filter α} {p : Prop} {q : α → Prop} : + (∃ᶠ x in f, p ∧ q x) ↔ p ∧ ∃ᶠ x in f, q x := by + simp only [Filter.Frequently, not_and, eventually_imp_distrib_left, Classical.not_imp] + +@[simp] +theorem frequently_and_distrib_right {f : Filter α} {p : α → Prop} {q : Prop} : + (∃ᶠ x in f, p x ∧ q) ↔ (∃ᶠ x in f, p x) ∧ q := by + simp only [@and_comm _ q, frequently_and_distrib_left] + +/-! +### Relation “eventually equal” +-/ + +section EventuallyEq +variable {l : Filter α} {f g : α → β} + +variable {l : Filter α} + +protected lemma EventuallyLE.iUnion [Finite ι] {s t : ι → Set α} + (h : ∀ i, s i ≤ᶠ[l] t i) : (⋃ i, s i) ≤ᶠ[l] ⋃ i, t i := + (eventually_all.2 h).mono fun _x hx hx' ↦ + let ⟨i, hi⟩ := mem_iUnion.1 hx'; mem_iUnion.2 ⟨i, hx i hi⟩ + +protected lemma EventuallyEq.iUnion [Finite ι] {s t : ι → Set α} + (h : ∀ i, s i =ᶠ[l] t i) : (⋃ i, s i) =ᶠ[l] ⋃ i, t i := + (EventuallyLE.iUnion fun i ↦ (h i).le).antisymm <| .iUnion fun i ↦ (h i).symm.le + +protected lemma EventuallyLE.iInter [Finite ι] {s t : ι → Set α} + (h : ∀ i, s i ≤ᶠ[l] t i) : (⋂ i, s i) ≤ᶠ[l] ⋂ i, t i := + (eventually_all.2 h).mono fun _x hx hx' ↦ mem_iInter.2 fun i ↦ hx i (mem_iInter.1 hx' i) + +protected lemma EventuallyEq.iInter [Finite ι] {s t : ι → Set α} + (h : ∀ i, s i =ᶠ[l] t i) : (⋂ i, s i) =ᶠ[l] ⋂ i, t i := + (EventuallyLE.iInter fun i ↦ (h i).le).antisymm <| .iInter fun i ↦ (h i).symm.le + +lemma _root_.Set.Finite.eventuallyLE_iUnion {ι : Type*} {s : Set ι} (hs : s.Finite) + {f g : ι → Set α} (hle : ∀ i ∈ s, f i ≤ᶠ[l] g i) : (⋃ i ∈ s, f i) ≤ᶠ[l] (⋃ i ∈ s, g i) := by + have := hs.to_subtype + rw [biUnion_eq_iUnion, biUnion_eq_iUnion] + exact .iUnion fun i ↦ hle i.1 i.2 + +alias EventuallyLE.biUnion := Set.Finite.eventuallyLE_iUnion + +lemma _root_.Set.Finite.eventuallyEq_iUnion {ι : Type*} {s : Set ι} (hs : s.Finite) + {f g : ι → Set α} (heq : ∀ i ∈ s, f i =ᶠ[l] g i) : (⋃ i ∈ s, f i) =ᶠ[l] (⋃ i ∈ s, g i) := + (EventuallyLE.biUnion hs fun i hi ↦ (heq i hi).le).antisymm <| + .biUnion hs fun i hi ↦ (heq i hi).symm.le + +alias EventuallyEq.biUnion := Set.Finite.eventuallyEq_iUnion + +lemma _root_.Set.Finite.eventuallyLE_iInter {ι : Type*} {s : Set ι} (hs : s.Finite) + {f g : ι → Set α} (hle : ∀ i ∈ s, f i ≤ᶠ[l] g i) : (⋂ i ∈ s, f i) ≤ᶠ[l] (⋂ i ∈ s, g i) := by + have := hs.to_subtype + rw [biInter_eq_iInter, biInter_eq_iInter] + exact .iInter fun i ↦ hle i.1 i.2 + +alias EventuallyLE.biInter := Set.Finite.eventuallyLE_iInter + +lemma _root_.Set.Finite.eventuallyEq_iInter {ι : Type*} {s : Set ι} (hs : s.Finite) + {f g : ι → Set α} (heq : ∀ i ∈ s, f i =ᶠ[l] g i) : (⋂ i ∈ s, f i) =ᶠ[l] (⋂ i ∈ s, g i) := + (EventuallyLE.biInter hs fun i hi ↦ (heq i hi).le).antisymm <| + .biInter hs fun i hi ↦ (heq i hi).symm.le + +alias EventuallyEq.biInter := Set.Finite.eventuallyEq_iInter + +lemma _root_.Finset.eventuallyLE_iUnion {ι : Type*} (s : Finset ι) {f g : ι → Set α} + (hle : ∀ i ∈ s, f i ≤ᶠ[l] g i) : (⋃ i ∈ s, f i) ≤ᶠ[l] (⋃ i ∈ s, g i) := + .biUnion s.finite_toSet hle + +lemma _root_.Finset.eventuallyEq_iUnion {ι : Type*} (s : Finset ι) {f g : ι → Set α} + (heq : ∀ i ∈ s, f i =ᶠ[l] g i) : (⋃ i ∈ s, f i) =ᶠ[l] (⋃ i ∈ s, g i) := + .biUnion s.finite_toSet heq + +lemma _root_.Finset.eventuallyLE_iInter {ι : Type*} (s : Finset ι) {f g : ι → Set α} + (hle : ∀ i ∈ s, f i ≤ᶠ[l] g i) : (⋂ i ∈ s, f i) ≤ᶠ[l] (⋂ i ∈ s, g i) := + .biInter s.finite_toSet hle + +lemma _root_.Finset.eventuallyEq_iInter {ι : Type*} (s : Finset ι) {f g : ι → Set α} + (heq : ∀ i ∈ s, f i =ᶠ[l] g i) : (⋂ i ∈ s, f i) =ᶠ[l] (⋂ i ∈ s, g i) := + .biInter s.finite_toSet heq + +end EventuallyEq + +end Filter + +open Filter diff --git a/Mathlib/Order/Filter/Germ/Basic.lean b/Mathlib/Order/Filter/Germ/Basic.lean index 00156fd9037ce..c9ad7754c4946 100644 --- a/Mathlib/Order/Filter/Germ/Basic.lean +++ b/Mathlib/Order/Filter/Germ/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Abhimanyu Pallavi Sudhir -/ import Mathlib.Algebra.Module.Pi +import Mathlib.Algebra.Order.Monoid.Unbundled.ExistsOfLE import Mathlib.Data.Int.Cast.Lemmas import Mathlib.Order.Filter.Tendsto diff --git a/Mathlib/Order/Filter/Subsingleton.lean b/Mathlib/Order/Filter/Subsingleton.lean index e079522700c04..c8d0816c3f17b 100644 --- a/Mathlib/Order/Filter/Subsingleton.lean +++ b/Mathlib/Order/Filter/Subsingleton.lean @@ -3,7 +3,6 @@ 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.Order.Filter.Bases import Mathlib.Order.Filter.Ultrafilter /-! # Subsingleton filters diff --git a/Mathlib/SetTheory/Cardinal/CountableCover.lean b/Mathlib/SetTheory/Cardinal/CountableCover.lean index 16eb4b46b329c..b3a0fe7e59ed2 100644 --- a/Mathlib/SetTheory/Cardinal/CountableCover.lean +++ b/Mathlib/SetTheory/Cardinal/CountableCover.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import Mathlib.SetTheory.Cardinal.Arithmetic -import Mathlib.Order.Filter.Basic +import Mathlib.Order.Filter.Finite /-! # Cardinality of a set with a countable cover From 56428e8ac3a752358608d5efe49acdf7da3bf885 Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 25 Nov 2024 09:04:46 +0000 Subject: [PATCH 269/829] fix(CI): checkout the script in the action (#19449) Previously, the script was failing, since the script was unavailable. --- .github/workflows/maintainer_bors.yml | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/.github/workflows/maintainer_bors.yml b/.github/workflows/maintainer_bors.yml index 290ea7c66fa1b..67e3596b6a67e 100644 --- a/.github/workflows/maintainer_bors.yml +++ b/.github/workflows/maintainer_bors.yml @@ -101,6 +101,12 @@ jobs: python -m pip install --upgrade pip pip install zulip + - uses: actions/checkout@v4 + with: + ref: master + sparse-checkout: | + scripts/zulip_emoji_merge_delegate.py + - name: Run Zulip Emoji Merge Delegate Script if: ${{ ! steps.merge_or_delegate.outputs.mOrD == '' && ( steps.user_permission.outputs.require-result == 'true' || From 31eae5146f7f5b2884a2a96e50b93caa2932cd61 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Mon, 25 Nov 2024 09:04:47 +0000 Subject: [PATCH 270/829] chore(discover-lean-pr-testing): only run on lean-toolchain changes (#19450) --- .github/workflows/discover-lean-pr-testing.yml | 2 ++ 1 file changed, 2 insertions(+) diff --git a/.github/workflows/discover-lean-pr-testing.yml b/.github/workflows/discover-lean-pr-testing.yml index 02e087dfb6085..349945a0ee8fe 100644 --- a/.github/workflows/discover-lean-pr-testing.yml +++ b/.github/workflows/discover-lean-pr-testing.yml @@ -4,6 +4,8 @@ on: push: branches: - nightly-testing + paths: + - lean-toolchain jobs: discover-lean-pr-testing: From 5fde9dcf705fe1d732a15da3dda14a624a95237f Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 25 Nov 2024 09:04:48 +0000 Subject: [PATCH 271/829] fix: escape line breaks in create adaptation PR (#19451) --- scripts/create-adaptation-pr.sh | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/scripts/create-adaptation-pr.sh b/scripts/create-adaptation-pr.sh index b8ae7cf37b28f..f8b5deb203b7e 100755 --- a/scripts/create-adaptation-pr.sh +++ b/scripts/create-adaptation-pr.sh @@ -198,7 +198,7 @@ if git diff --name-only bump/$BUMPVERSION bump/nightly-$NIGHTLYDATE | grep -q .; echo "### [auto] post a link to the PR on Zulip" zulip_title="#$pr_number adaptations for nightly-$NIGHTLYDATE" - zulip_body="> $pr_title #$pr_number\\\n\\\nPlease review this PR. At the end of the month this diff will land in \\\`master\\\`." + zulip_body="> $pr_title #$pr_number"$'\n\nPlease review this PR. At the end of the month this diff will land in `master`.' echo "Posting the link to the PR in a new thread on the #nightly-testing channel on Zulip" echo "Here is the message:" From 5f4575362b440d45be2cc96e0654bf03a6a0e479 Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Mon, 25 Nov 2024 09:33:01 +0000 Subject: [PATCH 272/829] refactor(Algebra/Category): turn homs in `AlgebraCat` into a structure (#19065) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In the current design of concrete categories in mathlib, there are two sources of `erw`s: - Def-Eq abuse of `A ⟶ B` and `A →ₐ [R] B`: The type of `A ⟶ B` is by definition `A →ₐ[R] B`, but not at reducible transparency. So it happens that one `rw` lemma transforms a term in a form, where the next `rw` expects `A →ₐ B`. By supplying explicit `show` lines (copying the output of the pretty-printer), Lean correctly re-synthesizes the type to be `A →ₐ B` and the `rw` succeeds. This def-eq abuse hence always causes issues in the case where a lemma from the non-category theoretic part of the library should be applied. An insightful example is the following proof excerpt from current master (proof of `AlgebraCat.adj`): ```lean import Mathlib.Algebra.Category.AlgebraCat.Basic open CategoryTheory AlgebraCat variable {R : Type u} [CommRing R] set_option pp.analyze true example : free.{u} R ⊣ forget (AlgebraCat.{u} R) := Adjunction.mkOfHomEquiv { homEquiv := fun _ _ => (FreeAlgebra.lift _).symm homEquiv_naturality_left_symm := by intro X Y Z f g apply FreeAlgebra.hom_ext ext x simp only [Equiv.symm_symm, Function.comp_apply, free_map] /- Eq (α := Z) (((FreeAlgebra.lift R) (f ≫ g) : Prefunctor.obj (Functor.toPrefunctor (free R)) X ⟶ Z) (FreeAlgebra.ι R x) : ↑Z) _ -/ erw [FreeAlgebra.lift_ι_apply] sorry homEquiv_naturality_right := sorry } ``` The `simp` lemma `FreeAlgebra.lift_ι_apply` expects an `AlgHom`, but as the `pp.analyze` shows, the type is synthesized as a `⟶`. With a show line before the `erw`, Lean correctly synthesizes the type as an `AlgHom` again and the `rw` succeeds. The solution is to strictly distinguish between `A ⟶ B` and `A →ₐ[R] B`: We define a new `AlgebraCat.Hom` structure with a single field `hom : A →ₐ[R] B`. The above proof is mostly solved by `ext; simp` then, except for the `ext` lemma `FreeAlgebra.hom_ext` not applying. This is because now the type is `AlgebraCat.of R (FreeAlgebra R X) →ₐ[R] _` and Lean can't see through the `AlgebraCat.of` at reducible transparency. Hence the solution is to make `AlgebraCat.of` an `abbrev`. Finally, the proof is `by ext; simp` or even `by aesop`. - `FunLike`: An intermediate design adapted the changes from the first point, but with keeping a `FunLike` and an `AlgHomClass` instance on `A ⟶ B`. This lead to many additional `coe` lemmas for composition of morphisms, where it mattered which of the three appearing terms of `AlgebraCat` where of the form `AlgebraCat.of R A`. Eric Wieser suggested to replace the `FunLike` by a `CoeFun` which is inlined, so an invocation of `f x` turns into `f.hom x`. This has then worked very smoothly. So the key points are: - Homs in concrete categories should be one-field structures to maintain a strict type distinction. - Use `CoeFun` instead of `FunLike`, since the latter is automatically inlined. - Make `.of` constructors an `abbrev` to obtain the def-eq `↑(AlgebraCat.of R A) = A` at reducible transparency. For more details, see the [Zulip discussion](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/Fixing.20the.20.60erw.60.20hell.20around.20concrete.20categories). --- .../Algebra/Category/AlgebraCat/Basic.lean | 218 +++++++++--------- .../Algebra/Category/AlgebraCat/Limits.lean | 52 ++--- .../Algebra/Category/AlgebraCat/Monoidal.lean | 2 +- .../Algebra/Category/BialgebraCat/Basic.lean | 4 +- .../Monoidal/Internal/Module.lean | 19 +- .../CliffordAlgebra/CategoryTheory.lean | 8 +- .../ConcreteCategory/AlgebraCat.lean | 53 +++++ 7 files changed, 195 insertions(+), 161 deletions(-) create mode 100644 MathlibTest/CategoryTheory/ConcreteCategory/AlgebraCat.lean diff --git a/Mathlib/Algebra/Category/AlgebraCat/Basic.lean b/Mathlib/Algebra/Category/AlgebraCat/Basic.lean index 9c06d5e619d85..1053dac09d757 100644 --- a/Mathlib/Algebra/Category/AlgebraCat/Basic.lean +++ b/Mathlib/Algebra/Category/AlgebraCat/Basic.lean @@ -16,7 +16,6 @@ with the forgetful functors to `RingCat` and `ModuleCat`. We furthermore show th associating to a type the free `R`-algebra on that type is left adjoint to the forgetful functor. -/ - open CategoryTheory Limits universe v u @@ -25,6 +24,8 @@ variable (R : Type u) [CommRing R] /-- The category of R-algebras and their morphisms. -/ structure AlgebraCat where + private mk :: + /-- The underlying type. -/ carrier : Type v [isRing : Ring carrier] [isAlgebra : Algebra R carrier] @@ -46,22 +47,97 @@ instance : CoeSort (AlgebraCat R) (Type v) := attribute [coe] AlgebraCat.carrier +/-- The object in the category of R-algebras associated to a type equipped with the appropriate +typeclasses. This is the preferred way to construct a term of `AlgebraCat R`. -/ +abbrev of (X : Type v) [Ring X] [Algebra R X] : AlgebraCat.{v} R := + ⟨X⟩ + +lemma coe_of (X : Type v) [Ring X] [Algebra R X] : (of R X : Type v) = X := + rfl + +variable {R} in +/-- The type of morphisms in `AlgebraCat R`. -/ +@[ext] +structure Hom (A B : AlgebraCat.{v} R) where + private mk :: + /-- The underlying algebra map. -/ + hom : A →ₐ[R] B + instance : Category (AlgebraCat.{v} R) where - Hom A B := A →ₐ[R] B - id A := AlgHom.id R A - comp f g := g.comp f + Hom A B := Hom A B + id A := ⟨AlgHom.id R A⟩ + comp f g := ⟨g.hom.comp f.hom⟩ + +instance {M N : AlgebraCat.{v} R} : CoeFun (M ⟶ N) (fun _ ↦ M → N) where + coe f := f.hom -instance {M N : AlgebraCat.{v} R} : FunLike (M ⟶ N) M N := - AlgHom.funLike +@[simp] +lemma hom_id {A : AlgebraCat.{v} R} : (𝟙 A : A ⟶ A).hom = AlgHom.id R A := rfl + +/- Provided for rewriting. -/ +lemma id_apply (A : AlgebraCat.{v} R) (a : A) : + (𝟙 A : A ⟶ A) a = a := by simp + +@[simp] +lemma hom_comp {A B C : AlgebraCat.{v} R} (f : A ⟶ B) (g : B ⟶ C) : + (f ≫ g).hom = g.hom.comp f.hom := rfl + +/- Provided for rewriting. -/ +lemma comp_apply {A B C : AlgebraCat.{v} R} (f : A ⟶ B) (g : B ⟶ C) (a : A) : + (f ≫ g) a = g (f a) := by simp + +@[ext] +lemma hom_ext {A B : AlgebraCat.{v} R} {f g : A ⟶ B} (hf : f.hom = g.hom) : f = g := + Hom.ext hf + +/-- Typecheck an `AlgHom` as a morphism in `AlgebraCat R`. -/ +abbrev ofHom {R : Type u} [CommRing R] {X Y : Type v} [Ring X] [Algebra R X] [Ring Y] [Algebra R Y] + (f : X →ₐ[R] Y) : of R X ⟶ of R Y := + ⟨f⟩ -instance {M N : AlgebraCat.{v} R} : AlgHomClass (M ⟶ N) R M N := - AlgHom.algHomClass +lemma hom_ofHom {R : Type u} [CommRing R] {X Y : Type v} [Ring X] [Algebra R X] [Ring Y] + [Algebra R Y] (f : X →ₐ[R] Y) : (ofHom f).hom = f := rfl + +@[simp] +lemma ofHom_hom {A B : AlgebraCat.{v} R} (f : A ⟶ B) : + ofHom (Hom.hom f) = f := rfl + +@[simp] +lemma ofHom_id {X : Type v} [Ring X] [Algebra R X] : ofHom (AlgHom.id R X) = 𝟙 (of R X) := rfl + +@[simp] +lemma ofHom_comp {X Y Z : Type v} [Ring X] [Ring Y] [Ring Z] [Algebra R X] [Algebra R Y] + [Algebra R Z] (f : X →ₐ[R] Y) (g : Y →ₐ[R] Z) : + ofHom (g.comp f) = ofHom f ≫ ofHom g := + rfl + +lemma ofHom_apply {R : Type u} [CommRing R] {X Y : Type v} [Ring X] [Algebra R X] [Ring Y] + [Algebra R Y] (f : X →ₐ[R] Y) (x : X) : ofHom f x = f x := rfl + +@[simp] +lemma inv_hom_apply {A B : AlgebraCat.{v} R} (e : A ≅ B) (x : A) : e.inv (e.hom x) = x := by + rw [← comp_apply] + simp + +@[simp] +lemma hom_inv_apply {A B : AlgebraCat.{v} R} (e : A ≅ B) (x : B) : e.hom (e.inv x) = x := by + rw [← comp_apply] + simp + +instance : Inhabited (AlgebraCat R) := + ⟨of R R⟩ instance : ConcreteCategory.{v} (AlgebraCat.{v} R) where forget := { obj := fun R => R - map := fun f => f.toFun } - forget_faithful := ⟨fun h => AlgHom.ext (by intros x; dsimp at h; rw [h])⟩ + map := fun f => f.hom } + forget_faithful := ⟨fun h => by ext x; simpa using congrFun h x⟩ + +lemma forget_obj {A : AlgebraCat.{v} R} : (forget (AlgebraCat.{v} R)).obj A = A := rfl + +lemma forget_map {A B : AlgebraCat.{v} R} (f : A ⟶ B) : + (forget (AlgebraCat.{v} R)).map f = f := + rfl instance {S : AlgebraCat.{v} R} : Ring ((forget (AlgebraCat R)).obj S) := (inferInstance : Ring S.carrier) @@ -72,12 +148,12 @@ instance {S : AlgebraCat.{v} R} : Algebra R ((forget (AlgebraCat R)).obj S) := instance hasForgetToRing : HasForget₂ (AlgebraCat.{v} R) RingCat.{v} where forget₂ := { obj := fun A => RingCat.of A - map := fun f => RingCat.ofHom f.toRingHom } + map := fun f => RingCat.ofHom f.hom.toRingHom } instance hasForgetToModule : HasForget₂ (AlgebraCat.{v} R) (ModuleCat.{v} R) where forget₂ := { obj := fun M => ModuleCat.of R M - map := fun f => ModuleCat.asHom f.toLinearMap } + map := fun f => ModuleCat.asHom f.hom.toLinearMap } @[simp] lemma forget₂_module_obj (X : AlgebraCat.{v} R) : @@ -86,33 +162,10 @@ lemma forget₂_module_obj (X : AlgebraCat.{v} R) : @[simp] lemma forget₂_module_map {X Y : AlgebraCat.{v} R} (f : X ⟶ Y) : - (forget₂ (AlgebraCat.{v} R) (ModuleCat.{v} R)).map f = ModuleCat.asHom f.toLinearMap := - rfl - -/-- The object in the category of R-algebras associated to a type equipped with the appropriate -typeclasses. -/ -def of (X : Type v) [Ring X] [Algebra R X] : AlgebraCat.{v} R := - ⟨X⟩ - -/-- Typecheck a `AlgHom` as a morphism in `AlgebraCat R`. -/ -def ofHom {R : Type u} [CommRing R] {X Y : Type v} [Ring X] [Algebra R X] [Ring Y] [Algebra R Y] - (f : X →ₐ[R] Y) : of R X ⟶ of R Y := - f - -@[simp] -theorem ofHom_apply {R : Type u} [CommRing R] {X Y : Type v} [Ring X] [Algebra R X] [Ring Y] - [Algebra R Y] (f : X →ₐ[R] Y) (x : X) : ofHom f x = f x := + (forget₂ (AlgebraCat.{v} R) (ModuleCat.{v} R)).map f = ModuleCat.asHom f.hom.toLinearMap := rfl -instance : Inhabited (AlgebraCat R) := - ⟨of R R⟩ - -@[simp] -theorem coe_of (X : Type u) [Ring X] [Algebra R X] : (of R X : Type u) = X := - rfl - -variable {R} - +variable {R} in /-- Forgetting to the underlying type and then building the bundled object returns the original algebra. -/ @[simps] @@ -120,58 +173,20 @@ def ofSelfIso (M : AlgebraCat.{v} R) : AlgebraCat.of R M ≅ M where hom := 𝟙 M inv := 𝟙 M -variable {M N U : ModuleCat.{v} R} - -@[simp] -theorem id_apply (m : M) : (𝟙 M : M → M) m = m := - rfl - -@[simp] -theorem coe_comp (f : M ⟶ N) (g : N ⟶ U) : (f ≫ g : M → U) = g ∘ f := - rfl - -variable (R) - /-- The "free algebra" functor, sending a type `S` to the free algebra on `S`. -/ -@[simps!] +@[simps! obj map] def free : Type u ⥤ AlgebraCat.{u} R where - obj S := - { carrier := FreeAlgebra R S - isRing := Algebra.semiringToRing R } - map f := FreeAlgebra.lift _ <| FreeAlgebra.ι _ ∘ f - -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11041): `apply FreeAlgebra.hom_ext` was `ext1`. - map_id := by intro X; apply FreeAlgebra.hom_ext; simp only [FreeAlgebra.ι_comp_lift]; rfl - map_comp := by - -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11041): `apply FreeAlgebra.hom_ext` was `ext1`. - intros; apply FreeAlgebra.hom_ext; simp only [FreeAlgebra.ι_comp_lift]; ext1 - -- Porting node: this ↓ `erw` used to be handled by the `simp` below it - erw [CategoryTheory.coe_comp] - simp only [CategoryTheory.coe_comp, Function.comp_apply, types_comp_apply] - -- Porting node: this ↓ `erw` and `rfl` used to be handled by the `simp` above - erw [FreeAlgebra.lift_ι_apply, FreeAlgebra.lift_ι_apply] - rfl + obj S := of R (FreeAlgebra R S) + map f := ofHom <| FreeAlgebra.lift _ <| FreeAlgebra.ι _ ∘ f /-- The free/forget adjunction for `R`-algebras. -/ def adj : free.{u} R ⊣ forget (AlgebraCat.{u} R) := Adjunction.mkOfHomEquiv - { homEquiv := fun _ _ => (FreeAlgebra.lift _).symm - -- Relying on `obviously` to fill out these proofs is very slow :( - homEquiv_naturality_left_symm := by - -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11041): `apply FreeAlgebra.hom_ext` was `ext1`. - intros; apply FreeAlgebra.hom_ext; simp only [FreeAlgebra.ι_comp_lift]; ext1 - simp only [free_map, Equiv.symm_symm, FreeAlgebra.lift_ι_apply, CategoryTheory.coe_comp, - Function.comp_apply, types_comp_apply] - -- Porting node: this ↓ `erw` and `rfl` used to be handled by the `simp` above - erw [FreeAlgebra.lift_ι_apply, CategoryTheory.comp_apply, FreeAlgebra.lift_ι_apply, - Function.comp_apply, FreeAlgebra.lift_ι_apply] - rfl - homEquiv_naturality_right := by - intros; ext - simp only [CategoryTheory.coe_comp, Function.comp_apply, - FreeAlgebra.lift_symm_apply, types_comp_apply] - -- Porting note: proof used to be done after this ↑ `simp`; added ↓ two lines - erw [FreeAlgebra.lift_symm_apply, FreeAlgebra.lift_symm_apply] - rfl } + { homEquiv := fun _ _ => + { toFun := fun f ↦ (FreeAlgebra.lift _).symm f.hom + invFun := fun f ↦ ofHom <| (FreeAlgebra.lift _) f + left_inv := fun f ↦ by aesop + right_inv := fun f ↦ by simp [forget_obj, forget_map] } } instance : (forget (AlgebraCat.{u} R)).IsRightAdjoint := (adj R).isRightAdjoint @@ -184,31 +199,19 @@ variable {X₁ X₂ : Type u} @[simps] def AlgEquiv.toAlgebraIso {g₁ : Ring X₁} {g₂ : Ring X₂} {m₁ : Algebra R X₁} {m₂ : Algebra R X₂} (e : X₁ ≃ₐ[R] X₂) : AlgebraCat.of R X₁ ≅ AlgebraCat.of R X₂ where - hom := (e : X₁ →ₐ[R] X₂) - inv := (e.symm : X₂ →ₐ[R] X₁) - hom_inv_id := by ext x; exact e.left_inv x - inv_hom_id := by ext x; exact e.right_inv x + hom := AlgebraCat.ofHom (e : X₁ →ₐ[R] X₂) + inv := AlgebraCat.ofHom (e.symm : X₂ →ₐ[R] X₁) namespace CategoryTheory.Iso /-- Build a `AlgEquiv` from an isomorphism in the category `AlgebraCat R`. -/ @[simps] def toAlgEquiv {X Y : AlgebraCat R} (i : X ≅ Y) : X ≃ₐ[R] Y := - { i.hom with + { i.hom.hom with toFun := i.hom invFun := i.inv - left_inv := fun x => by - -- Porting note: was `by tidy` - change (i.hom ≫ i.inv) x = x - simp only [hom_inv_id] - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [id_apply] - right_inv := fun x => by - -- Porting note: was `by tidy` - change (i.inv ≫ i.hom) x = x - simp only [inv_hom_id] - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [id_apply] } + left_inv := fun x ↦ by simp + right_inv := fun x ↦ by simp } end CategoryTheory.Iso @@ -223,18 +226,5 @@ def algEquivIsoAlgebraIso {X Y : Type u} [Ring X] [Ring Y] [Algebra R X] [Algebr instance AlgebraCat.forget_reflects_isos : (forget (AlgebraCat.{u} R)).ReflectsIsomorphisms where reflects {X Y} f _ := by let i := asIso ((forget (AlgebraCat.{u} R)).map f) - let e : X ≃ₐ[R] Y := { f, i.toEquiv with } + let e : X ≃ₐ[R] Y := { f.hom, i.toEquiv with } exact e.toAlgebraIso.isIso_hom - -/-! -`@[simp]` lemmas for `AlgHom.comp` and categorical identities. --/ - -@[simp] theorem AlgHom.comp_id_algebraCat - {R} [CommRing R] {G : AlgebraCat.{u} R} {H : Type u} [Ring H] [Algebra R H] (f : G →ₐ[R] H) : - f.comp (𝟙 G) = f := - Category.id_comp (AlgebraCat.ofHom f) -@[simp] theorem AlgHom.id_algebraCat_comp - {R} [CommRing R] {G : Type u} [Ring G] [Algebra R G] {H : AlgebraCat.{u} R} (f : G →ₐ[R] H) : - AlgHom.comp (𝟙 H) f = f := - Category.comp_id (AlgebraCat.ofHom f) diff --git a/Mathlib/Algebra/Category/AlgebraCat/Limits.lean b/Mathlib/Algebra/Category/AlgebraCat/Limits.lean index 94f312fee5673..ea39419a2442e 100644 --- a/Mathlib/Algebra/Category/AlgebraCat/Limits.lean +++ b/Mathlib/Algebra/Category/AlgebraCat/Limits.lean @@ -40,7 +40,7 @@ instance algebraObj (j) : def sectionsSubalgebra : Subalgebra R (∀ j, F.obj j) := { SemiRingCat.sectionsSubsemiring (F ⋙ forget₂ (AlgebraCat R) RingCat.{w} ⋙ forget₂ RingCat SemiRingCat.{w}) with - algebraMap_mem' := fun r _ _ f => (F.map f).commutes r } + algebraMap_mem' := fun r _ _ f => (F.map f).hom.commutes r } instance (F : J ⥤ AlgebraCat.{w} R) : Ring (F ⋙ forget _).sections := inferInstanceAs <| Ring (sectionsSubalgebra F) @@ -88,9 +88,10 @@ namespace HasLimits def limitCone : Cone F where pt := AlgebraCat.of R (Types.Small.limitCone (F ⋙ forget _)).pt π := - { app := limitπAlgHom F - naturality := fun _ _ f => - AlgHom.coe_fn_injective ((Types.Small.limitCone (F ⋙ forget _)).π.naturality f) } + { app := fun j ↦ ofHom <| limitπAlgHom F j + naturality := fun _ _ f => by + ext : 1 + exact AlgHom.coe_fn_injective ((Types.Small.limitCone (F ⋙ forget _)).π.naturality f) } /-- Witness that the limit cone in `AlgebraCat R` is a limit cone. (Internal use only; use the limits API.) @@ -101,41 +102,36 @@ def limitConeIsLimit : IsLimit (limitCone.{v, w} F) := by -- Porting note: in mathlib3 the function term -- `fun v => ⟨fun j => ((forget (AlgebraCat R)).mapCone s).π.app j v` -- was provided by unification, and the last argument `(fun s => _)` was `(fun s => rfl)`. - (fun s => { toFun := _, map_one' := ?_, map_mul' := ?_, map_zero' := ?_, map_add' := ?_, - commutes' := ?_ }) + (fun s => ofHom + { toFun := _, map_one' := ?_, map_mul' := ?_, map_zero' := ?_, map_add' := ?_, + commutes' := ?_ }) (fun s => rfl) · congr ext j - simp only [Functor.comp_obj, Functor.mapCone_pt, Functor.mapCone_π_app, - forget_map_eq_coe] - rw [map_one] - rfl + simp only [Functor.mapCone_π_app, forget_map, map_one, Pi.one_apply] · intro x y - simp only [Functor.comp_obj, Functor.mapCone_pt, Functor.mapCone_π_app] - rw [← equivShrink_mul] - apply congrArg ext j - simp only [Functor.comp_obj, Functor.mapCone_pt, Functor.mapCone_π_app, - forget_map_eq_coe, map_mul] - rfl - · simp only [Functor.mapCone_π_app, forget_map_eq_coe] - congr - funext j - simp only [map_zero, Pi.zero_apply] + simp only [Functor.comp_obj, forget_obj, Equiv.toFun_as_coe, Functor.mapCone_pt, + Functor.mapCone_π_app, forget_map, Equiv.symm_apply_apply, + Types.Small.limitCone_pt, equivShrink_symm_mul] + apply map_mul + · ext j + simp only [Functor.comp_obj, forget_obj, Equiv.toFun_as_coe, Functor.mapCone_pt, + Functor.mapCone_π_app, forget_map, Equiv.symm_apply_apply, + equivShrink_symm_zero] + apply map_zero · intro x y - simp only [Functor.mapCone_π_app] - rw [← equivShrink_add] - apply congrArg ext j - simp only [forget_map_eq_coe, map_add] - rfl + simp only [Functor.comp_obj, forget_obj, Equiv.toFun_as_coe, Functor.mapCone_pt, + Functor.mapCone_π_app, forget_map, Equiv.symm_apply_apply, + Types.Small.limitCone_pt, equivShrink_symm_add] + apply map_add · intro r - simp only [← Shrink.algEquiv_symm_apply _ R, limitCone, Equiv.algebraMap_def, - Equiv.symm_symm] + simp only [← Shrink.algEquiv_symm_apply _ R, limitCone, Equiv.algebraMap_def, Equiv.symm_symm] apply congrArg apply Subtype.ext ext j - exact (s.π.app j).commutes r + exact (s.π.app j).hom.commutes r end HasLimits diff --git a/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean b/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean index ac5df5ca0f2dd..fa16a8aac4bcd 100644 --- a/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean +++ b/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean @@ -37,7 +37,7 @@ noncomputable abbrev tensorObj (X Y : AlgebraCat.{u} R) : AlgebraCat.{u} R := `AlgebraCat.instMonoidalCategory`. -/ noncomputable abbrev tensorHom {W X Y Z : AlgebraCat.{u} R} (f : W ⟶ X) (g : Y ⟶ Z) : tensorObj W Y ⟶ tensorObj X Z := - Algebra.TensorProduct.map f g + ofHom <| Algebra.TensorProduct.map f.hom g.hom open MonoidalCategory diff --git a/Mathlib/Algebra/Category/BialgebraCat/Basic.lean b/Mathlib/Algebra/Category/BialgebraCat/Basic.lean index fb1f40ffe1018..59fe5b8fefba2 100644 --- a/Mathlib/Algebra/Category/BialgebraCat/Basic.lean +++ b/Mathlib/Algebra/Category/BialgebraCat/Basic.lean @@ -103,7 +103,7 @@ instance concreteCategory : ConcreteCategory.{v} (BialgebraCat.{v} R) where instance hasForgetToAlgebra : HasForget₂ (BialgebraCat R) (AlgebraCat R) where forget₂ := { obj := fun X => AlgebraCat.of R X - map := fun {X Y} f => (f.toBialgHom : X →ₐ[R] Y) } + map := fun {X Y} f => AlgebraCat.ofHom f.toBialgHom } @[simp] theorem forget₂_algebra_obj (X : BialgebraCat R) : @@ -112,7 +112,7 @@ theorem forget₂_algebra_obj (X : BialgebraCat R) : @[simp] theorem forget₂_algebra_map (X Y : BialgebraCat R) (f : X ⟶ Y) : - (forget₂ (BialgebraCat R) (AlgebraCat R)).map f = (f.toBialgHom : X →ₐ[R] Y) := + (forget₂ (BialgebraCat R) (AlgebraCat R)).map f = AlgebraCat.ofHom f.toBialgHom := rfl instance hasForgetToCoalgebra : HasForget₂ (BialgebraCat R) (CoalgebraCat R) where diff --git a/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean b/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean index 5974753934292..9341953c58518 100644 --- a/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean +++ b/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean @@ -88,7 +88,7 @@ theorem algebraMap (A : Mon_ (ModuleCat.{u} R)) (r : R) : algebraMap R A.X r = A @[simps!] def functor : Mon_ (ModuleCat.{u} R) ⥤ AlgebraCat R where obj A := AlgebraCat.of R A.X - map {_ _} f := + map {_ _} f := AlgebraCat.ofHom { f.hom.toAddMonoidHom with toFun := f.hom map_one' := LinearMap.congr_fun f.one_hom (1 : R) @@ -134,9 +134,7 @@ def inverseObj (A : AlgebraCat.{u} R) : Mon_ (ModuleCat.{u} R) where -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): `ext` did not pick up `TensorProduct.ext` refine TensorProduct.ext <| TensorProduct.ext <| LinearMap.ext fun x => LinearMap.ext fun y => LinearMap.ext fun z => ?_ - dsimp only [AlgebraCat.id_apply, TensorProduct.mk_apply, LinearMap.compr₂_apply, - Function.comp_apply, ModuleCat.MonoidalCategory.tensorHom_tmul, AlgebraCat.coe_comp, - MonoidalCategory.associator_hom_apply] + dsimp only [compr₂_apply, TensorProduct.mk_apply] rw [compr₂_apply, compr₂_apply] -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 erw [CategoryTheory.comp_apply, @@ -152,9 +150,9 @@ def inverseObj (A : AlgebraCat.{u} R) : Mon_ (ModuleCat.{u} R) where def inverse : AlgebraCat.{u} R ⥤ Mon_ (ModuleCat.{u} R) where obj := inverseObj map f := - { hom := f.toLinearMap - one_hom := LinearMap.ext f.commutes - mul_hom := TensorProduct.ext <| LinearMap.ext₂ <| map_mul f } + { hom := f.hom.toLinearMap + one_hom := LinearMap.ext f.hom.commutes + mul_hom := TensorProduct.ext <| LinearMap.ext₂ <| map_mul f.hom } end MonModuleEquivalenceAlgebra @@ -193,14 +191,14 @@ def monModuleEquivalenceAlgebra : Mon_ (ModuleCat.{u} R) ≌ AlgebraCat R where counitIso := NatIso.ofComponents (fun A => - { hom := + { hom := AlgebraCat.ofHom { toFun := _root_.id map_zero' := rfl map_add' := fun _ _ => rfl map_one' := (algebraMap R A).map_one map_mul' := fun x y => @LinearMap.mul'_apply R _ _ _ _ _ _ x y commutes' := fun _ => rfl } - inv := + inv := AlgebraCat.ofHom { toFun := _root_.id map_zero' := rfl map_add' := fun _ _ => rfl @@ -208,9 +206,6 @@ def monModuleEquivalenceAlgebra : Mon_ (ModuleCat.{u} R) ≌ AlgebraCat R where map_mul' := fun x y => (@LinearMap.mul'_apply R _ _ _ _ _ _ x y).symm commutes' := fun _ => rfl } }) --- These lemmas have always been bad (https://github.com/leanprover-community/mathlib4/issues/7657), but https://github.com/leanprover/lean4/pull/2644 made `simp` start noticing -attribute [nolint simpNF] ModuleCat.MonModuleEquivalenceAlgebra.functor_map_apply - /-- The equivalence `Mon_ (ModuleCat R) ≌ AlgebraCat R` is naturally compatible with the forgetful functors to `ModuleCat R`. -/ diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/CategoryTheory.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/CategoryTheory.lean index 2311b81349cc4..0b5239188f120 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/CategoryTheory.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/CategoryTheory.lean @@ -25,7 +25,7 @@ variable {R : Type u} [CommRing R] This is `CliffordAlgebra.map` through the lens of category theory. -/ @[simps] def QuadraticModuleCat.cliffordAlgebra : QuadraticModuleCat.{u} R ⥤ AlgebraCat.{u} R where - obj M := { carrier := CliffordAlgebra M.form } - map {_M _N} f := CliffordAlgebra.map f.toIsometry - map_id _M := CliffordAlgebra.map_id _ - map_comp {_M _N _P} f g := (CliffordAlgebra.map_comp_map g.toIsometry f.toIsometry).symm + obj M := AlgebraCat.of R (CliffordAlgebra M.form) + map {_M _N} f := AlgebraCat.ofHom <| CliffordAlgebra.map f.toIsometry + map_id _M := by simp + map_comp {_M _N _P} f g := by ext; simp diff --git a/MathlibTest/CategoryTheory/ConcreteCategory/AlgebraCat.lean b/MathlibTest/CategoryTheory/ConcreteCategory/AlgebraCat.lean new file mode 100644 index 0000000000000..406d12cc89de4 --- /dev/null +++ b/MathlibTest/CategoryTheory/ConcreteCategory/AlgebraCat.lean @@ -0,0 +1,53 @@ +import Mathlib + +universe v u + +open CategoryTheory AlgebraCat + +set_option maxHeartbeats 10000 +set_option synthInstance.maxHeartbeats 2000 + +variable (R : Type u) [CommRing R] + +/- We test if all the coercions and `map_add` lemmas trigger correctly. -/ + +example (X : Type u) [Ring X] [Algebra R X] : ⇑(𝟙 (of R X)) = id := by simp + +example {X Y : Type v} [Ring X] [Algebra R X] [Ring Y] [Algebra R Y] (f : X →ₐ[R] Y) : + ⇑(ofHom f) = ⇑f := by simp + +example {X Y : Type v} [Ring X] [Algebra R X] [Ring Y] [Algebra R Y] (f : X →ₐ[R] Y) + (x : X) : (ofHom f) x = f x := by simp + +example {X Y Z : AlgebraCat R} (f : X ⟶ Y) (g : Y ⟶ Z) : ⇑(f ≫ g) = ⇑g ∘ ⇑f := by simp + +example {X Y Z : Type v} [Ring X] [Algebra R X] [Ring Y] [Algebra R Y] [Ring Z] + [Algebra R Z] (f : X →ₐ[R] Y) (g : Y →ₐ[R] Z) : + ⇑(ofHom f ≫ ofHom g) = g ∘ f := by simp + +example {X Y : Type v} [Ring X] [Algebra R X] [Ring Y] [Algebra R Y] {Z : AlgebraCat R} + (f : X →ₐ[R] Y) (g : of R Y ⟶ Z) : + ⇑(ofHom f ≫ g) = g ∘ f := by simp + +example {X Y : AlgebraCat R} {Z : Type v} [Ring Z] [Algebra R Z] (f : X ⟶ Y) (g : Y ⟶ of R Z) : + ⇑(f ≫ g) = g ∘ f := by simp + +example {Y Z : AlgebraCat R} {X : Type v} [Ring X] [Algebra R X] (f : of R X ⟶ Y) (g : Y ⟶ Z) : + ⇑(f ≫ g) = g ∘ f := by simp + +example {X Y Z : AlgebraCat R} (f : X ⟶ Y) (g : Y ⟶ Z) (x : X) : (f ≫ g) x = g (f x) := by simp + +example {X Y : AlgebraCat R} (e : X ≅ Y) (x : X) : e.inv (e.hom x) = x := by simp + +example {X Y : AlgebraCat R} (e : X ≅ Y) (y : Y) : e.hom (e.inv y) = y := by simp + +example (X : AlgebraCat R) : ⇑(𝟙 X) = id := by simp + +example {M N : AlgebraCat.{v} R} (f : M ⟶ N) (x y : M) : f (x + y) = f x + f y := by + simp + +example {M N : AlgebraCat.{v} R} (f : M ⟶ N) : f 0 = 0 := by + simp + +example {M N : AlgebraCat.{v} R} (f : M ⟶ N) (r : R) (m : M) : f (r • m) = r • f m := by + simp From a5ae0c1c33f439193d3e3e967e514035edb9fe55 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Mon, 25 Nov 2024 09:33:02 +0000 Subject: [PATCH 273/829] chore: remove the `Smooth` alias for infinitely differentiable functions between manifolds (#19296) The abbreviation `Smooth` for `ContMDiff` with infinite smoothness had been introduced in mathlib in ancient times. This has lead to a big API duplication, for no visible gain. We deprecate this abbreviation (with its friends `SmoothAt` and so on) and all the lemmas involving it. In the longer run, we might consider reintroducing it as a mere notation (without any specific API), but this is not part of this PR, devoted purely to the deprecation. No mathematical content added. --- .../Complex/UpperHalfPlane/Manifold.lean | 16 +- .../Distribution/AEEqOfIntegralContDiff.lean | 10 +- .../Geometry/Manifold/Algebra/LieGroup.lean | 125 ++++------- Mathlib/Geometry/Manifold/Algebra/Monoid.lean | 177 +++++++--------- .../Manifold/Algebra/SmoothFunctions.lean | 30 +-- .../Geometry/Manifold/Algebra/Structures.lean | 8 +- Mathlib/Geometry/Manifold/BumpFunction.lean | 21 +- .../Geometry/Manifold/ContMDiff/Atlas.lean | 4 +- .../Geometry/Manifold/ContMDiff/Basic.lean | 114 ++++------ Mathlib/Geometry/Manifold/ContMDiff/Defs.lean | 135 ++++-------- .../Manifold/ContMDiff/NormedSpace.lean | 27 +-- .../Geometry/Manifold/ContMDiff/Product.lean | 124 ++++------- .../Geometry/Manifold/ContMDiffMFDeriv.lean | 11 +- Mathlib/Geometry/Manifold/ContMDiffMap.lean | 8 +- Mathlib/Geometry/Manifold/Diffeomorph.lean | 12 +- Mathlib/Geometry/Manifold/MFDeriv/Basic.lean | 29 +-- .../Manifold/MFDeriv/UniqueDifferential.lean | 2 +- .../Geometry/Manifold/PartitionOfUnity.lean | 77 ++++--- .../Manifold/Sheaf/LocallyRingedSpace.lean | 2 +- Mathlib/Geometry/Manifold/Sheaf/Smooth.lean | 11 +- .../Geometry/Manifold/VectorBundle/Basic.lean | 198 ++++++++---------- .../VectorBundle/FiberwiseLinear.lean | 44 ++-- .../Geometry/Manifold/VectorBundle/Hom.lean | 21 +- .../Manifold/VectorBundle/Pullback.lean | 6 +- .../Manifold/VectorBundle/SmoothSection.lean | 10 +- .../Manifold/VectorBundle/Tangent.lean | 2 +- .../Geometry/Manifold/WhitneyEmbedding.lean | 13 +- 27 files changed, 495 insertions(+), 742 deletions(-) diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Manifold.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Manifold.lean index c821c8feafdcb..9529dbcb764f2 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Manifold.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Manifold.lean @@ -26,15 +26,17 @@ instance : SmoothManifoldWithCorners 𝓘(ℂ) ℍ := UpperHalfPlane.isOpenEmbedding_coe.singleton_smoothManifoldWithCorners /-- The inclusion map `ℍ → ℂ` is a smooth map of manifolds. -/ -theorem smooth_coe : Smooth 𝓘(ℂ) 𝓘(ℂ) ((↑) : ℍ → ℂ) := fun _ => contMDiffAt_extChartAt +theorem contMDiff_coe : ContMDiff 𝓘(ℂ) 𝓘(ℂ) ⊤ ((↑) : ℍ → ℂ) := fun _ => contMDiffAt_extChartAt + +@[deprecated (since := "2024-11-20")] alias smooth_coe := contMDiff_coe /-- The inclusion map `ℍ → ℂ` is a differentiable map of manifolds. -/ theorem mdifferentiable_coe : MDifferentiable 𝓘(ℂ) 𝓘(ℂ) ((↑) : ℍ → ℂ) := - smooth_coe.mdifferentiable + contMDiff_coe.mdifferentiable (by simp) -lemma smoothAt_ofComplex {z : ℂ} (hz : 0 < z.im) : - SmoothAt 𝓘(ℂ) 𝓘(ℂ) ofComplex z := by - rw [SmoothAt, contMDiffAt_iff] +lemma contMDiffAt_ofComplex {z : ℂ} (hz : 0 < z.im) : + ContMDiffAt 𝓘(ℂ) 𝓘(ℂ) ⊤ ofComplex z := by + rw [contMDiffAt_iff] constructor · -- continuity at z rw [ContinuousAt, nhds_induced, tendsto_comap_iff] @@ -48,9 +50,11 @@ lemma smoothAt_ofComplex {z : ℂ} (hz : 0 < z.im) : Set.range_id, id_eq, contDiffWithinAt_univ] exact contDiffAt_id.congr_of_eventuallyEq (eventuallyEq_coe_comp_ofComplex hz) +@[deprecated (since := "2024-11-20")] alias smoothAt_ofComplex := contMDiffAt_ofComplex + lemma mdifferentiableAt_ofComplex {z : ℂ} (hz : 0 < z.im) : MDifferentiableAt 𝓘(ℂ) 𝓘(ℂ) ofComplex z := - (smoothAt_ofComplex hz).mdifferentiableAt + (contMDiffAt_ofComplex hz).mdifferentiableAt (by simp) lemma mdifferentiableAt_iff {f : ℍ → ℂ} {τ : ℍ} : MDifferentiableAt 𝓘(ℂ) 𝓘(ℂ) f τ ↔ DifferentiableAt ℂ (f ∘ ofComplex) ↑τ := by diff --git a/Mathlib/Analysis/Distribution/AEEqOfIntegralContDiff.lean b/Mathlib/Analysis/Distribution/AEEqOfIntegralContDiff.lean index fb4a816b271bc..b6d2c67574d34 100644 --- a/Mathlib/Analysis/Distribution/AEEqOfIntegralContDiff.lean +++ b/Mathlib/Analysis/Distribution/AEEqOfIntegralContDiff.lean @@ -40,7 +40,7 @@ variable {H : Type*} [TopologicalSpace H] (I : ModelWithCorners ℝ E H) when multiplied by any smooth compactly supported function, then `f` vanishes almost everywhere. -/ theorem ae_eq_zero_of_integral_smooth_smul_eq_zero [SigmaCompactSpace M] (hf : LocallyIntegrable f μ) - (h : ∀ g : M → ℝ, Smooth I 𝓘(ℝ) g → HasCompactSupport g → ∫ x, g x • f x ∂μ = 0) : + (h : ∀ g : M → ℝ, ContMDiff I 𝓘(ℝ) ⊤ g → HasCompactSupport g → ∫ x, g x • f x ∂μ = 0) : ∀ᵐ x ∂μ, f x = 0 := by -- record topological properties of `M` have := I.locallyCompactSpace @@ -60,7 +60,7 @@ theorem ae_eq_zero_of_integral_smooth_smul_eq_zero [SigmaCompactSpace M] let v : ℕ → Set M := fun n ↦ thickening (u n) s obtain ⟨K, K_compact, vK⟩ : ∃ K, IsCompact K ∧ ∀ n, v n ⊆ K := ⟨_, hδ, fun n ↦ thickening_subset_cthickening_of_le (u_pos n).2.le _⟩ - have : ∀ n, ∃ (g : M → ℝ), support g = v n ∧ Smooth I 𝓘(ℝ) g ∧ Set.range g ⊆ Set.Icc 0 1 + have : ∀ n, ∃ (g : M → ℝ), support g = v n ∧ ContMDiff I 𝓘(ℝ) ⊤ g ∧ Set.range g ⊆ Set.Icc 0 1 ∧ ∀ x ∈ s, g x = 1 := by intro n rcases exists_msmooth_support_eq_eq_one_iff I isOpen_thickening hs.isClosed @@ -121,7 +121,7 @@ instance (U : Opens M) : BorelSpace U := inferInstanceAs (BorelSpace (U : Set M) nonrec theorem IsOpen.ae_eq_zero_of_integral_smooth_smul_eq_zero' {U : Set M} (hU : IsOpen U) (hSig : IsSigmaCompact U) (hf : LocallyIntegrableOn f U μ) (h : ∀ g : M → ℝ, - Smooth I 𝓘(ℝ) g → HasCompactSupport g → tsupport g ⊆ U → ∫ x, g x • f x ∂μ = 0) : + ContMDiff I 𝓘(ℝ) ⊤ g → HasCompactSupport g → tsupport g ⊆ U → ∫ x, g x • f x ∂μ = 0) : ∀ᵐ x ∂μ, x ∈ U → f x = 0 := by have meas_U := hU.measurableSet rw [← ae_restrict_iff' meas_U, ae_restrict_iff_subtype meas_U] @@ -146,7 +146,7 @@ variable [SigmaCompactSpace M] theorem IsOpen.ae_eq_zero_of_integral_smooth_smul_eq_zero {U : Set M} (hU : IsOpen U) (hf : LocallyIntegrableOn f U μ) (h : ∀ g : M → ℝ, - Smooth I 𝓘(ℝ) g → HasCompactSupport g → tsupport g ⊆ U → ∫ x, g x • f x ∂μ = 0) : + ContMDiff I 𝓘(ℝ) ⊤ g → HasCompactSupport g → tsupport g ⊆ U → ∫ x, g x • f x ∂μ = 0) : ∀ᵐ x ∂μ, x ∈ U → f x = 0 := haveI := I.locallyCompactSpace haveI := ChartedSpace.locallyCompactSpace H M @@ -160,7 +160,7 @@ theorem IsOpen.ae_eq_zero_of_integral_smooth_smul_eq_zero {U : Set M} (hU : IsOp when multiplied by any smooth compactly supported function, then they coincide almost everywhere. -/ theorem ae_eq_of_integral_smooth_smul_eq (hf : LocallyIntegrable f μ) (hf' : LocallyIntegrable f' μ) (h : ∀ (g : M → ℝ), - Smooth I 𝓘(ℝ) g → HasCompactSupport g → ∫ x, g x • f x ∂μ = ∫ x, g x • f' x ∂μ) : + ContMDiff I 𝓘(ℝ) ⊤ g → HasCompactSupport g → ∫ x, g x • f x ∂μ = ∫ x, g x • f' x ∂μ) : ∀ᵐ x ∂μ, f x = f' x := by have : ∀ᵐ x ∂μ, (f - f') x = 0 := by apply ae_eq_zero_of_integral_smooth_smul_eq_zero I (hf.sub hf') diff --git a/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean b/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean index c054be8efa60b..c9dd661d44afd 100644 --- a/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean +++ b/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean @@ -58,7 +58,7 @@ class LieAddGroup {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [Top {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] (I : ModelWithCorners 𝕜 E H) (G : Type*) [AddGroup G] [TopologicalSpace G] [ChartedSpace H G] extends SmoothAdd I G : Prop where /-- Negation is smooth in an additive Lie group. -/ - smooth_neg : Smooth I I fun a : G => -a + smooth_neg : ContMDiff I I ⊤ fun a : G => -a -- See note [Design choices about smooth algebraic structures] /-- A (multiplicative) Lie group is a group and a smooth manifold at the same time in which @@ -68,7 +68,7 @@ class LieGroup {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [Topolo {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] (I : ModelWithCorners 𝕜 E H) (G : Type*) [Group G] [TopologicalSpace G] [ChartedSpace H G] extends SmoothMul I G : Prop where /-- Inversion is smooth in a Lie group. -/ - smooth_inv : Smooth I I fun a : G => a⁻¹ + smooth_inv : ContMDiff I I ⊤ fun a : G => a⁻¹ /-! ### Smoothness of inversion, negation, division and subtraction @@ -91,28 +91,31 @@ variable (I) /-- In a Lie group, inversion is a smooth map. -/ @[to_additive "In an additive Lie group, inversion is a smooth map."] -theorem smooth_inv : Smooth I I fun x : G => x⁻¹ := +theorem contMDiff_inv : ContMDiff I I ⊤ fun x : G => x⁻¹ := LieGroup.smooth_inv +@[deprecated (since := "2024-11-21")] alias smooth_inv := contMDiff_inv +@[deprecated (since := "2024-11-21")] alias smooth_neg := contMDiff_neg + include I in /-- A Lie group is a topological group. This is not an instance for technical reasons, see note [Design choices about smooth algebraic structures]. -/ @[to_additive "An additive Lie group is an additive topological group. This is not an instance for technical reasons, see note [Design choices about smooth algebraic structures]."] theorem topologicalGroup_of_lieGroup : TopologicalGroup G := - { continuousMul_of_smooth I with continuous_inv := (smooth_inv I).continuous } + { continuousMul_of_smooth I with continuous_inv := (contMDiff_inv I).continuous } end @[to_additive] theorem ContMDiffWithinAt.inv {f : M → G} {s : Set M} {x₀ : M} (hf : ContMDiffWithinAt I' I n f s x₀) : ContMDiffWithinAt I' I n (fun x => (f x)⁻¹) s x₀ := - ((smooth_inv I).of_le le_top).contMDiffAt.contMDiffWithinAt.comp x₀ hf <| Set.mapsTo_univ _ _ + ((contMDiff_inv I).of_le le_top).contMDiffAt.contMDiffWithinAt.comp x₀ hf <| Set.mapsTo_univ _ _ @[to_additive] theorem ContMDiffAt.inv {f : M → G} {x₀ : M} (hf : ContMDiffAt I' I n f x₀) : ContMDiffAt I' I n (fun x => (f x)⁻¹) x₀ := - ((smooth_inv I).of_le le_top).contMDiffAt.comp x₀ hf + ((contMDiff_inv I).of_le le_top).contMDiffAt.comp x₀ hf @[to_additive] theorem ContMDiffOn.inv {f : M → G} {s : Set M} (hf : ContMDiffOn I' I n f s) : @@ -122,24 +125,15 @@ theorem ContMDiffOn.inv {f : M → G} {s : Set M} (hf : ContMDiffOn I' I n f s) theorem ContMDiff.inv {f : M → G} (hf : ContMDiff I' I n f) : ContMDiff I' I n fun x => (f x)⁻¹ := fun x => (hf x).inv -@[to_additive] -nonrec theorem SmoothWithinAt.inv {f : M → G} {s : Set M} {x₀ : M} - (hf : SmoothWithinAt I' I f s x₀) : SmoothWithinAt I' I (fun x => (f x)⁻¹) s x₀ := - hf.inv - -@[to_additive] -nonrec theorem SmoothAt.inv {f : M → G} {x₀ : M} (hf : SmoothAt I' I f x₀) : - SmoothAt I' I (fun x => (f x)⁻¹) x₀ := - hf.inv +@[deprecated (since := "2024-11-21")] alias SmoothWithinAt.inv := ContMDiffWithinAt.inv +@[deprecated (since := "2024-11-21")] alias SmoothAt.inv := ContMDiffAt.inv +@[deprecated (since := "2024-11-21")] alias SmoothOn.inv := ContMDiffOn.inv +@[deprecated (since := "2024-11-21")] alias Smooth.inv := ContMDiff.inv -@[to_additive] -nonrec theorem SmoothOn.inv {f : M → G} {s : Set M} (hf : SmoothOn I' I f s) : - SmoothOn I' I (fun x => (f x)⁻¹) s := - hf.inv - -@[to_additive] -nonrec theorem Smooth.inv {f : M → G} (hf : Smooth I' I f) : Smooth I' I fun x => (f x)⁻¹ := - hf.inv +@[deprecated (since := "2024-11-21")] alias SmoothWithinAt.neg := ContMDiffWithinAt.neg +@[deprecated (since := "2024-11-21")] alias SmoothAt.neg := ContMDiffAt.neg +@[deprecated (since := "2024-11-21")] alias SmoothOn.neg := ContMDiffOn.neg +@[deprecated (since := "2024-11-21")] alias Smooth.neg := ContMDiff.neg @[to_additive] theorem ContMDiffWithinAt.div {f g : M → G} {s : Set M} {x₀ : M} @@ -161,26 +155,15 @@ theorem ContMDiffOn.div {f g : M → G} {s : Set M} (hf : ContMDiffOn I' I n f s theorem ContMDiff.div {f g : M → G} (hf : ContMDiff I' I n f) (hg : ContMDiff I' I n g) : ContMDiff I' I n fun x => f x / g x := by simp_rw [div_eq_mul_inv]; exact hf.mul hg.inv -@[to_additive] -nonrec theorem SmoothWithinAt.div {f g : M → G} {s : Set M} {x₀ : M} - (hf : SmoothWithinAt I' I f s x₀) (hg : SmoothWithinAt I' I g s x₀) : - SmoothWithinAt I' I (fun x => f x / g x) s x₀ := - hf.div hg +@[deprecated (since := "2024-11-21")] alias SmoothWithinAt.div := ContMDiffWithinAt.div +@[deprecated (since := "2024-11-21")] alias SmoothAt.div := ContMDiffAt.div +@[deprecated (since := "2024-11-21")] alias SmoothOn.div := ContMDiffOn.div +@[deprecated (since := "2024-11-21")] alias Smooth.div := ContMDiff.div -@[to_additive] -nonrec theorem SmoothAt.div {f g : M → G} {x₀ : M} (hf : SmoothAt I' I f x₀) - (hg : SmoothAt I' I g x₀) : SmoothAt I' I (fun x => f x / g x) x₀ := - hf.div hg - -@[to_additive] -nonrec theorem SmoothOn.div {f g : M → G} {s : Set M} (hf : SmoothOn I' I f s) - (hg : SmoothOn I' I g s) : SmoothOn I' I (f / g) s := - hf.div hg - -@[to_additive] -nonrec theorem Smooth.div {f g : M → G} (hf : Smooth I' I f) (hg : Smooth I' I g) : - Smooth I' I (f / g) := - hf.div hg +@[deprecated (since := "2024-11-21")] alias SmoothWithinAt.sub := ContMDiffWithinAt.sub +@[deprecated (since := "2024-11-21")] alias SmoothAt.sub := ContMDiffAt.sub +@[deprecated (since := "2024-11-21")] alias SmoothOn.sub := ContMDiffOn.sub +@[deprecated (since := "2024-11-21")] alias Smooth.sub := ContMDiff.sub end PointwiseDivision @@ -195,7 +178,7 @@ instance {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [TopologicalS [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} {G' : Type*} [TopologicalSpace G'] [ChartedSpace H' G'] [Group G'] [LieGroup I' G'] : LieGroup (I.prod I') (G × G') := - { SmoothMul.prod _ _ _ _ with smooth_inv := smooth_fst.inv.prod_mk smooth_snd.inv } + { SmoothMul.prod _ _ _ _ with smooth_inv := contMDiff_fst.inv.prod_mk contMDiff_snd.inv } end Product @@ -219,7 +202,7 @@ class SmoothInv₀ {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [To {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] (I : ModelWithCorners 𝕜 E H) (G : Type*) [Inv G] [Zero G] [TopologicalSpace G] [ChartedSpace H G] : Prop where /-- Inversion is smooth away from `0`. -/ - smoothAt_inv₀ : ∀ ⦃x : G⦄, x ≠ 0 → SmoothAt I I (fun y ↦ y⁻¹) x + smoothAt_inv₀ : ∀ ⦃x : G⦄, x ≠ 0 → ContMDiffAt I I ⊤ (fun y ↦ y⁻¹) x instance {𝕜 : Type*} [NontriviallyNormedField 𝕜] [CompleteSpace 𝕜] : SmoothInv₀ 𝓘(𝕜) 𝕜 := { smoothAt_inv₀ := by @@ -235,28 +218,33 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [TopologicalS {I' : ModelWithCorners 𝕜 E' H'} {M : Type*} [TopologicalSpace M] [ChartedSpace H' M] {n : ℕ∞} {f : M → G} -theorem smoothAt_inv₀ {x : G} (hx : x ≠ 0) : SmoothAt I I (fun y ↦ y⁻¹) x := +theorem contMDiffAt_inv₀ {x : G} (hx : x ≠ 0) : ContMDiffAt I I ⊤ (fun y ↦ y⁻¹) x := SmoothInv₀.smoothAt_inv₀ hx +@[deprecated (since := "2024-11-21")] alias smoothAt_inv₀ := contMDiffAt_inv₀ + include I in /-- In a manifold with smooth inverse away from `0`, the inverse is continuous away from `0`. This is not an instance for technical reasons, see note [Design choices about smooth algebraic structures]. -/ theorem hasContinuousInv₀_of_hasSmoothInv₀ : HasContinuousInv₀ G := - { continuousAt_inv₀ := fun _ hx ↦ (smoothAt_inv₀ I hx).continuousAt } + { continuousAt_inv₀ := fun _ hx ↦ (contMDiffAt_inv₀ I hx).continuousAt } -theorem SmoothOn_inv₀ : SmoothOn I I (Inv.inv : G → G) {0}ᶜ := fun _x hx => - (smoothAt_inv₀ I hx).smoothWithinAt +theorem contMDiffOn_inv₀ : ContMDiffOn I I ⊤ (Inv.inv : G → G) {0}ᶜ := fun _x hx => + (contMDiffAt_inv₀ I hx).contMDiffWithinAt + +@[deprecated (since := "2024-11-21")] alias smoothOn_inv₀ := contMDiffOn_inv₀ +@[deprecated (since := "2024-11-21")] alias SmoothOn_inv₀ := contMDiffOn_inv₀ variable {I} {s : Set M} {a : M} theorem ContMDiffWithinAt.inv₀ (hf : ContMDiffWithinAt I' I n f s a) (ha : f a ≠ 0) : ContMDiffWithinAt I' I n (fun x => (f x)⁻¹) s a := - (smoothAt_inv₀ I ha).contMDiffAt.comp_contMDiffWithinAt a hf + ((contMDiffAt_inv₀ I ha).of_le le_top).comp_contMDiffWithinAt a hf theorem ContMDiffAt.inv₀ (hf : ContMDiffAt I' I n f a) (ha : f a ≠ 0) : ContMDiffAt I' I n (fun x ↦ (f x)⁻¹) a := - (smoothAt_inv₀ I ha).contMDiffAt.comp a hf + ((contMDiffAt_inv₀ I ha).of_le le_top).comp a hf theorem ContMDiff.inv₀ (hf : ContMDiff I' I n f) (h0 : ∀ x, f x ≠ 0) : ContMDiff I' I n (fun x ↦ (f x)⁻¹) := @@ -266,20 +254,10 @@ theorem ContMDiffOn.inv₀ (hf : ContMDiffOn I' I n f s) (h0 : ∀ x ∈ s, f x ContMDiffOn I' I n (fun x => (f x)⁻¹) s := fun x hx ↦ ContMDiffWithinAt.inv₀ (hf x hx) (h0 x hx) -theorem SmoothWithinAt.inv₀ (hf : SmoothWithinAt I' I f s a) (ha : f a ≠ 0) : - SmoothWithinAt I' I (fun x => (f x)⁻¹) s a := - ContMDiffWithinAt.inv₀ hf ha - -theorem SmoothAt.inv₀ (hf : SmoothAt I' I f a) (ha : f a ≠ 0) : - SmoothAt I' I (fun x => (f x)⁻¹) a := - ContMDiffAt.inv₀ hf ha - -theorem Smooth.inv₀ (hf : Smooth I' I f) (h0 : ∀ x, f x ≠ 0) : Smooth I' I fun x => (f x)⁻¹ := - ContMDiff.inv₀ hf h0 - -theorem SmoothOn.inv₀ (hf : SmoothOn I' I f s) (h0 : ∀ x ∈ s, f x ≠ 0) : - SmoothOn I' I (fun x => (f x)⁻¹) s := - ContMDiffOn.inv₀ hf h0 +@[deprecated (since := "2024-11-21")] alias SmoothWithinAt.inv₀ := ContMDiffWithinAt.inv₀ +@[deprecated (since := "2024-11-21")] alias SmoothAt.inv₀ := ContMDiffAt.inv₀ +@[deprecated (since := "2024-11-21")] alias SmoothOn.inv₀ := ContMDiffOn.inv₀ +@[deprecated (since := "2024-11-21")] alias Smooth.inv₀ := ContMDiff.inv₀ end SmoothInv₀ @@ -313,20 +291,9 @@ theorem ContMDiffAt.div₀ (hf : ContMDiffAt I' I n f a) (hg : ContMDiffAt I' I theorem ContMDiff.div₀ (hf : ContMDiff I' I n f) (hg : ContMDiff I' I n g) (h₀ : ∀ x, g x ≠ 0) : ContMDiff I' I n (f / g) := by simpa only [div_eq_mul_inv] using hf.mul (hg.inv₀ h₀) -theorem SmoothWithinAt.div₀ (hf : SmoothWithinAt I' I f s a) - (hg : SmoothWithinAt I' I g s a) (h₀ : g a ≠ 0) : SmoothWithinAt I' I (f / g) s a := - ContMDiffWithinAt.div₀ hf hg h₀ - -theorem SmoothOn.div₀ (hf : SmoothOn I' I f s) (hg : SmoothOn I' I g s) (h₀ : ∀ x ∈ s, g x ≠ 0) : - SmoothOn I' I (f / g) s := - ContMDiffOn.div₀ hf hg h₀ - -theorem SmoothAt.div₀ (hf : SmoothAt I' I f a) (hg : SmoothAt I' I g a) (h₀ : g a ≠ 0) : - SmoothAt I' I (f / g) a := - ContMDiffAt.div₀ hf hg h₀ - -theorem Smooth.div₀ (hf : Smooth I' I f) (hg : Smooth I' I g) (h₀ : ∀ x, g x ≠ 0) : - Smooth I' I (f / g) := - ContMDiff.div₀ hf hg h₀ +@[deprecated (since := "2024-11-21")] alias SmoothWithinAt.div₀ := ContMDiffWithinAt.div₀ +@[deprecated (since := "2024-11-21")] alias SmoothAt.div₀ := ContMDiffAt.div₀ +@[deprecated (since := "2024-11-21")] alias SmoothOn.div₀ := ContMDiffOn.div₀ +@[deprecated (since := "2024-11-21")] alias Smooth.div₀ := ContMDiff.div₀ end Div diff --git a/Mathlib/Geometry/Manifold/Algebra/Monoid.lean b/Mathlib/Geometry/Manifold/Algebra/Monoid.lean index 5439b39e461ed..3808ac40a51c1 100644 --- a/Mathlib/Geometry/Manifold/Algebra/Monoid.lean +++ b/Mathlib/Geometry/Manifold/Algebra/Monoid.lean @@ -45,7 +45,7 @@ class SmoothAdd {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [Topol {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] (I : ModelWithCorners 𝕜 E H) (G : Type*) [Add G] [TopologicalSpace G] [ChartedSpace H G] extends SmoothManifoldWithCorners I G : Prop where - smooth_add : Smooth (I.prod I) I fun p : G × G => p.1 + p.2 + smooth_add : ContMDiff (I.prod I) I ⊤ fun p : G × G => p.1 + p.2 -- See note [Design choices about smooth algebraic structures] /-- Basic hypothesis to talk about a smooth (Lie) monoid or a smooth semigroup. @@ -56,7 +56,7 @@ class SmoothMul {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [Topol {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] (I : ModelWithCorners 𝕜 E H) (G : Type*) [Mul G] [TopologicalSpace G] [ChartedSpace H G] extends SmoothManifoldWithCorners I G : Prop where - smooth_mul : Smooth (I.prod I) I fun p : G × G => p.1 * p.2 + smooth_mul : ContMDiff (I.prod I) I ⊤ fun p : G × G => p.1 * p.2 section SmoothMul @@ -71,16 +71,19 @@ section variable (I) @[to_additive] -theorem smooth_mul : Smooth (I.prod I) I fun p : G × G => p.1 * p.2 := +theorem contMDiff_mul : ContMDiff (I.prod I) I ⊤ fun p : G × G => p.1 * p.2 := SmoothMul.smooth_mul +@[deprecated (since := "2024-11-20")] alias smooth_mul := contMDiff_mul +@[deprecated (since := "2024-11-20")] alias smooth_add := contMDiff_add + include I in /-- If the multiplication is smooth, then it is continuous. This is not an instance for technical reasons, see note [Design choices about smooth algebraic structures]. -/ @[to_additive "If the addition is smooth, then it is continuous. This is not an instance for technical reasons, see note [Design choices about smooth algebraic structures]."] theorem continuousMul_of_smooth : ContinuousMul G := - ⟨(smooth_mul I).continuous⟩ + ⟨(contMDiff_mul I).continuous⟩ end @@ -91,7 +94,7 @@ variable {f g : M → G} {s : Set M} {x : M} {n : ℕ∞} @[to_additive] theorem ContMDiffWithinAt.mul (hf : ContMDiffWithinAt I' I n f s x) (hg : ContMDiffWithinAt I' I n g s x) : ContMDiffWithinAt I' I n (f * g) s x := - ((smooth_mul I).smoothAt.of_le le_top).comp_contMDiffWithinAt x (hf.prod_mk hg) + ((contMDiff_mul I).contMDiffAt.of_le le_top).comp_contMDiffWithinAt x (hf.prod_mk hg) @[to_additive] nonrec theorem ContMDiffAt.mul (hf : ContMDiffAt I' I n f x) (hg : ContMDiffAt I' I n g x) : @@ -106,52 +109,51 @@ theorem ContMDiffOn.mul (hf : ContMDiffOn I' I n f s) (hg : ContMDiffOn I' I n g theorem ContMDiff.mul (hf : ContMDiff I' I n f) (hg : ContMDiff I' I n g) : ContMDiff I' I n (f * g) := fun x => (hf x).mul (hg x) -@[to_additive] -nonrec theorem SmoothWithinAt.mul (hf : SmoothWithinAt I' I f s x) - (hg : SmoothWithinAt I' I g s x) : SmoothWithinAt I' I (f * g) s x := - hf.mul hg - -@[to_additive] -nonrec theorem SmoothAt.mul (hf : SmoothAt I' I f x) (hg : SmoothAt I' I g x) : - SmoothAt I' I (f * g) x := - hf.mul hg +@[deprecated (since := "2024-11-21")] alias SmoothWithinAt.mul := ContMDiffWithinAt.mul +@[deprecated (since := "2024-11-21")] alias SmoothAt.mul := ContMDiffAt.mul +@[deprecated (since := "2024-11-21")] alias SmoothOn.mul := ContMDiffOn.mul +@[deprecated (since := "2024-11-21")] alias Smooth.mul := ContMDiff.mul -@[to_additive] -nonrec theorem SmoothOn.mul (hf : SmoothOn I' I f s) (hg : SmoothOn I' I g s) : - SmoothOn I' I (f * g) s := - hf.mul hg +@[deprecated (since := "2024-11-21")] alias SmoothWithinAt.add := ContMDiffWithinAt.add +@[deprecated (since := "2024-11-21")] alias SmoothAt.add := ContMDiffAt.add +@[deprecated (since := "2024-11-21")] alias SmoothOn.add := ContMDiffOn.add +@[deprecated (since := "2024-11-21")] alias Smooth.add := ContMDiff.add @[to_additive] -nonrec theorem Smooth.mul (hf : Smooth I' I f) (hg : Smooth I' I g) : Smooth I' I (f * g) := - hf.mul hg +theorem contMDiff_mul_left {a : G} : ContMDiff I I n (a * ·) := + contMDiff_const.mul contMDiff_id -@[to_additive] -theorem smooth_mul_left {a : G} : Smooth I I fun b : G => a * b := - smooth_const.mul smooth_id +@[deprecated (since := "2024-11-21")] alias smooth_mul_left := contMDiff_mul_left +@[deprecated (since := "2024-11-21")] alias smooth_add_left := contMDiff_add_left @[to_additive] -theorem smooth_mul_right {a : G} : Smooth I I fun b : G => b * a := - smooth_id.mul smooth_const - -theorem contMDiff_mul_left {a : G} : ContMDiff I I n (a * ·) := smooth_mul_left.contMDiff - theorem contMDiffAt_mul_left {a b : G} : ContMDiffAt I I n (a * ·) b := contMDiff_mul_left.contMDiffAt +@[to_additive] theorem mdifferentiable_mul_left {a : G} : MDifferentiable I I (a * ·) := contMDiff_mul_left.mdifferentiable le_rfl +@[to_additive] theorem mdifferentiableAt_mul_left {a b : G} : MDifferentiableAt I I (a * ·) b := contMDiffAt_mul_left.mdifferentiableAt le_rfl -theorem contMDiff_mul_right {a : G} : ContMDiff I I n (· * a) := smooth_mul_right.contMDiff +@[to_additive] +theorem contMDiff_mul_right {a : G} : ContMDiff I I n (· * a) := + contMDiff_id.mul contMDiff_const + +@[deprecated (since := "2024-11-21")] alias smooth_mul_right := contMDiff_mul_right +@[deprecated (since := "2024-11-21")] alias smooth_add_right := contMDiff_add_right +@[to_additive] theorem contMDiffAt_mul_right {a b : G} : ContMDiffAt I I n (· * a) b := contMDiff_mul_right.contMDiffAt +@[to_additive] theorem mdifferentiable_mul_right {a : G} : MDifferentiable I I (· * a) := contMDiff_mul_right.mdifferentiable le_rfl +@[to_additive] theorem mdifferentiableAt_mul_right {a b : G} : MDifferentiableAt I I (· * a) b := contMDiffAt_mul_right.mdifferentiableAt le_rfl @@ -163,13 +165,13 @@ variable (I) (g h : G) Lemmas involving `smoothLeftMul` with the notation `𝑳` usually use `L` instead of `𝑳` in the names. -/ def smoothLeftMul : C^∞⟮I, G; I, G⟯ := - ⟨leftMul g, smooth_mul_left⟩ + ⟨leftMul g, contMDiff_mul_left⟩ /-- Right multiplication by `g`. It is meant to mimic the usual notation in Lie groups. Lemmas involving `smoothRightMul` with the notation `𝑹` usually use `R` instead of `𝑹` in the names. -/ def smoothRightMul : C^∞⟮I, G; I, G⟯ := - ⟨rightMul g, smooth_mul_right⟩ + ⟨rightMul g, contMDiff_mul_right⟩ -- Left multiplication. The abbreviation is `MIL`. scoped[LieGroup] notation "𝑳" => smoothLeftMul @@ -222,8 +224,8 @@ instance SmoothMul.prod {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type* [ChartedSpace H' G'] [Mul G'] [SmoothMul I' G'] : SmoothMul (I.prod I') (G × G') := { SmoothManifoldWithCorners.prod G G' with smooth_mul := - ((smooth_fst.comp smooth_fst).smooth.mul (smooth_fst.comp smooth_snd)).prod_mk - ((smooth_snd.comp smooth_fst).smooth.mul (smooth_snd.comp smooth_snd)) } + ((contMDiff_fst.comp contMDiff_fst).mul (contMDiff_fst.comp contMDiff_snd)).prod_mk + ((contMDiff_snd.comp contMDiff_fst).mul (contMDiff_snd.comp contMDiff_snd)) } end SmoothMul @@ -236,16 +238,19 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [TopologicalS {G' : Type*} [Monoid G'] [TopologicalSpace G'] [ChartedSpace H' G'] [SmoothMul I' G'] @[to_additive] -theorem smooth_pow : ∀ n : ℕ, Smooth I I fun a : G => a ^ n - | 0 => by simp only [pow_zero]; exact smooth_const - | k + 1 => by simpa [pow_succ] using (smooth_pow _).mul smooth_id +theorem contMDiff_pow : ∀ n : ℕ, ContMDiff I I ⊤ fun a : G => a ^ n + | 0 => by simp only [pow_zero]; exact contMDiff_const + | k + 1 => by simpa [pow_succ] using (contMDiff_pow _).mul contMDiff_id + +@[deprecated (since := "2024-11-21")] alias smooth_pow := contMDiff_pow +@[deprecated (since := "2024-11-21")] alias smooth_nsmul := contMDiff_nsmul /-- Morphism of additive smooth monoids. -/ structure SmoothAddMonoidMorphism (I : ModelWithCorners 𝕜 E H) (I' : ModelWithCorners 𝕜 E' H') (G : Type*) [TopologicalSpace G] [ChartedSpace H G] [AddMonoid G] [SmoothAdd I G] (G' : Type*) [TopologicalSpace G'] [ChartedSpace H' G'] [AddMonoid G'] [SmoothAdd I' G'] extends G →+ G' where - smooth_toFun : Smooth I I' toFun + smooth_toFun : ContMDiff I I' ⊤ toFun /-- Morphism of smooth monoids. -/ @[to_additive] @@ -253,11 +258,11 @@ structure SmoothMonoidMorphism (I : ModelWithCorners 𝕜 E H) (I' : ModelWithCo (G : Type*) [TopologicalSpace G] [ChartedSpace H G] [Monoid G] [SmoothMul I G] (G' : Type*) [TopologicalSpace G'] [ChartedSpace H' G'] [Monoid G'] [SmoothMul I' G'] extends G →* G' where - smooth_toFun : Smooth I I' toFun + smooth_toFun : ContMDiff I I' ⊤ toFun @[to_additive] instance : One (SmoothMonoidMorphism I I' G G') := - ⟨{ smooth_toFun := smooth_const + ⟨{ smooth_toFun := contMDiff_const toMonoidHom := 1 }⟩ @[to_additive] @@ -389,61 +394,43 @@ theorem contMDiff_finprod_cond (hc : ∀ i, p i → ContMDiff I' I n (f i)) simp only [← finprod_subtype_eq_finprod_cond] exact contMDiff_finprod (fun i => hc i i.2) (hf.comp_injective Subtype.coe_injective) -@[to_additive] -theorem smoothAt_finprod - (lf : LocallyFinite fun i ↦ mulSupport <| f i) (h : ∀ i, SmoothAt I' I (f i) x₀) : - SmoothAt I' I (fun x ↦ ∏ᶠ i, f i x) x₀ := - contMDiffWithinAt_finprod lf h +@[deprecated (since := "2024-11-21")] alias smoothAt_finprod := contMDiffAt_finprod +@[deprecated (since := "2024-11-21")] alias smoothAt_finsum := contMDiffAt_finsum -@[to_additive] -theorem smoothWithinAt_finset_prod' (h : ∀ i ∈ t, SmoothWithinAt I' I (f i) s x) : - SmoothWithinAt I' I (∏ i ∈ t, f i) s x := - contMDiffWithinAt_finset_prod' h +@[deprecated (since := "2024-11-21")] +alias smoothWithinAt_finset_prod' := contMDiffWithinAt_finset_prod' +@[deprecated (since := "2024-11-21")] +alias smoothWithinAt_finset_sum' := contMDiffWithinAt_finset_sum' -@[to_additive] -theorem smoothWithinAt_finset_prod (h : ∀ i ∈ t, SmoothWithinAt I' I (f i) s x) : - SmoothWithinAt I' I (fun x => ∏ i ∈ t, f i x) s x := - contMDiffWithinAt_finset_prod h -@[to_additive] -theorem smoothAt_finset_prod' (h : ∀ i ∈ t, SmoothAt I' I (f i) x) : - SmoothAt I' I (∏ i ∈ t, f i) x := - contMDiffAt_finset_prod' h +@[deprecated (since := "2024-11-21")] +alias smoothWithinAt_finset_prod := contMDiffWithinAt_finset_prod +@[deprecated (since := "2024-11-21")] +alias smoothWithinAt_finset_sum := contMDiffWithinAt_finset_sum -@[to_additive] -theorem smoothAt_finset_prod (h : ∀ i ∈ t, SmoothAt I' I (f i) x) : - SmoothAt I' I (fun x => ∏ i ∈ t, f i x) x := - contMDiffAt_finset_prod h +@[deprecated (since := "2024-11-21")] alias smoothAt_finset_prod' := contMDiffAt_finset_prod' +@[deprecated (since := "2024-11-21")] alias smoothAt_finset_sum' := contMDiffAt_finset_sum' -@[to_additive] -theorem smoothOn_finset_prod' (h : ∀ i ∈ t, SmoothOn I' I (f i) s) : - SmoothOn I' I (∏ i ∈ t, f i) s := - contMDiffOn_finset_prod' h +@[deprecated (since := "2024-11-21")] alias smoothAt_finset_prod := contMDiffAt_finset_prod +@[deprecated (since := "2024-11-21")] alias smoothAt_finset_sum := contMDiffAt_finset_sum -@[to_additive] -theorem smoothOn_finset_prod (h : ∀ i ∈ t, SmoothOn I' I (f i) s) : - SmoothOn I' I (fun x => ∏ i ∈ t, f i x) s := - contMDiffOn_finset_prod h +@[deprecated (since := "2024-11-21")] alias smoothOn_finset_prod' := contMDiffOn_finset_prod' +@[deprecated (since := "2024-11-21")] alias smoothOn_finset_sum' := contMDiffOn_finset_sum' -@[to_additive] -theorem smooth_finset_prod' (h : ∀ i ∈ t, Smooth I' I (f i)) : Smooth I' I (∏ i ∈ t, f i) := - contMDiff_finset_prod' h +@[deprecated (since := "2024-11-21")] alias smoothOn_finset_prod := contMDiffOn_finset_prod +@[deprecated (since := "2024-11-21")] alias smoothOn_finset_sum := contMDiffOn_finset_sum -@[to_additive] -theorem smooth_finset_prod (h : ∀ i ∈ t, Smooth I' I (f i)) : - Smooth I' I fun x => ∏ i ∈ t, f i x := - contMDiff_finset_prod h +@[deprecated (since := "2024-11-21")] alias smooth_finset_prod' := contMDiffOn_finset_prod' +@[deprecated (since := "2024-11-21")] alias smooth_finset_sum' := contMDiffOn_finset_sum' -@[to_additive] -theorem smooth_finprod (h : ∀ i, Smooth I' I (f i)) - (hfin : LocallyFinite fun i => mulSupport (f i)) : Smooth I' I fun x => ∏ᶠ i, f i x := - contMDiff_finprod h hfin +@[deprecated (since := "2024-11-21")] alias smooth_finset_prod := contMDiff_finset_prod +@[deprecated (since := "2024-11-21")] alias smooth_finset_sum := contMDiff_finset_sum -@[to_additive] -theorem smooth_finprod_cond (hc : ∀ i, p i → Smooth I' I (f i)) - (hf : LocallyFinite fun i => mulSupport (f i)) : - Smooth I' I fun x => ∏ᶠ (i) (_ : p i), f i x := - contMDiff_finprod_cond hc hf +@[deprecated (since := "2024-11-21")] alias smooth_finprod := contMDiff_finprod +@[deprecated (since := "2024-11-21")] alias smooth_finsum := contMDiff_finsum + +@[deprecated (since := "2024-11-21")] alias smooth_finprod_cond := contMDiff_finprod_cond +@[deprecated (since := "2024-11-21")] alias smooth_finsum_cond := contMDiff_finsum_cond end CommMonoid @@ -488,23 +475,9 @@ theorem ContMDiffOn.div_const (hf : ContMDiffOn I' I n f s) : theorem ContMDiff.div_const (hf : ContMDiff I' I n f) : ContMDiff I' I n (fun x ↦ f x / c) := fun x => (hf x).div_const c -@[to_additive] -nonrec theorem SmoothWithinAt.div_const (hf : SmoothWithinAt I' I f s x) : - SmoothWithinAt I' I (fun x ↦ f x / c) s x := - hf.div_const c - -@[to_additive] -nonrec theorem SmoothAt.div_const (hf : SmoothAt I' I f x) : - SmoothAt I' I (fun x ↦ f x / c) x := - hf.div_const c - -@[to_additive] -nonrec theorem SmoothOn.div_const (hf : SmoothOn I' I f s) : - SmoothOn I' I (fun x ↦ f x / c) s := - hf.div_const c - -@[to_additive] -nonrec theorem Smooth.div_const (hf : Smooth I' I f) : Smooth I' I (fun x ↦ f x / c) := - hf.div_const c +@[deprecated (since := "2024-11-21")] alias SmoothWithinAt.div_const := ContMDiffWithinAt.div_const +@[deprecated (since := "2024-11-21")] alias SmoothAt.div_const := ContMDiffAt.div_const +@[deprecated (since := "2024-11-21")] alias SmoothOn.div_const := ContMDiffOn.div_const +@[deprecated (since := "2024-11-21")] alias Smooth.div_const := ContMDiff.div_const end DivConst diff --git a/Mathlib/Geometry/Manifold/Algebra/SmoothFunctions.lean b/Mathlib/Geometry/Manifold/Algebra/SmoothFunctions.lean index 002ad1cf560cf..e8adb09c03d94 100644 --- a/Mathlib/Geometry/Manifold/Algebra/SmoothFunctions.lean +++ b/Mathlib/Geometry/Manifold/Algebra/SmoothFunctions.lean @@ -30,7 +30,7 @@ namespace SmoothMap @[to_additive] protected instance instMul {G : Type*} [Mul G] [TopologicalSpace G] [ChartedSpace H' G] [SmoothMul I' G] : Mul C^∞⟮I, N; I', G⟯ := - ⟨fun f g => ⟨f * g, f.smooth.mul g.smooth⟩⟩ + ⟨fun f g => ⟨f * g, f.contMDiff.mul g.contMDiff⟩⟩ @[to_additive (attr := simp)] theorem coe_mul {G : Type*} [Mul G] [TopologicalSpace G] [ChartedSpace H' G] [SmoothMul I' G] @@ -54,12 +54,12 @@ theorem coe_one {G : Type*} [One G] [TopologicalSpace G] [ChartedSpace H' G] : instance instNSMul {G : Type*} [AddMonoid G] [TopologicalSpace G] [ChartedSpace H' G] [SmoothAdd I' G] : SMul ℕ C^∞⟮I, N; I', G⟯ where - smul n f := ⟨n • (f : N → G), (smooth_nsmul n).comp f.smooth⟩ + smul n f := ⟨n • (f : N → G), (contMDiff_nsmul n).comp f.contMDiff⟩ @[to_additive existing] instance instPow {G : Type*} [Monoid G] [TopologicalSpace G] [ChartedSpace H' G] [SmoothMul I' G] : Pow C^∞⟮I, N; I', G⟯ ℕ where - pow f n := ⟨(f : N → G) ^ n, (smooth_pow n).comp f.smooth⟩ + pow f n := ⟨(f : N → G) ^ n, (contMDiff_pow n).comp f.contMDiff⟩ @[to_additive (attr := simp)] theorem coe_pow {G : Type*} [Monoid G] [TopologicalSpace G] [ChartedSpace H' G] [SmoothMul I' G] @@ -104,9 +104,9 @@ variable (I N) `C^∞⟮I, N; I'', G''⟯`."] def compLeftMonoidHom {G' : Type*} [Monoid G'] [TopologicalSpace G'] [ChartedSpace H' G'] [SmoothMul I' G'] {G'' : Type*} [Monoid G''] [TopologicalSpace G''] [ChartedSpace H'' G''] - [SmoothMul I'' G''] (φ : G' →* G'') (hφ : Smooth I' I'' φ) : + [SmoothMul I'' G''] (φ : G' →* G'') (hφ : ContMDiff I' I'' ⊤ φ) : C^∞⟮I, N; I', G'⟯ →* C^∞⟮I, N; I'', G''⟯ where - toFun f := ⟨φ ∘ f, fun x => (hφ.smooth _).comp x (f.contMDiff x)⟩ + toFun f := ⟨φ ∘ f, hφ.comp f.contMDiff⟩ map_one' := by ext; show φ 1 = 1; simp map_mul' f g := by ext x; show φ (f x * g x) = φ (f x) * φ (g x); simp @@ -119,7 +119,7 @@ variable (I') {N} homomorphism from `C^∞⟮I, V; I', G⟯` to `C^∞⟮I, U; I', G⟯`."] def restrictMonoidHom (G : Type*) [Monoid G] [TopologicalSpace G] [ChartedSpace H' G] [SmoothMul I' G] {U V : Opens N} (h : U ≤ V) : C^∞⟮I, V; I', G⟯ →* C^∞⟮I, U; I', G⟯ where - toFun f := ⟨f ∘ Set.inclusion h, f.smooth.comp (smooth_inclusion h)⟩ + toFun f := ⟨f ∘ Set.inclusion h, f.contMDiff.comp (contMDiff_inclusion h)⟩ map_one' := rfl map_mul' _ _ := rfl @@ -134,9 +134,9 @@ instance commMonoid {G : Type*} [CommMonoid G] [TopologicalSpace G] [ChartedSpac instance group {G : Type*} [Group G] [TopologicalSpace G] [ChartedSpace H' G] [LieGroup I' G] : Group C^∞⟮I, N; I', G⟯ := { SmoothMap.monoid with - inv := fun f => ⟨fun x => (f x)⁻¹, f.smooth.inv⟩ + inv := fun f => ⟨fun x => (f x)⁻¹, f.contMDiff.inv⟩ inv_mul_cancel := fun a => by ext; exact inv_mul_cancel _ - div := fun f g => ⟨f / g, f.smooth.div g.smooth⟩ + div := fun f g => ⟨f / g, f.contMDiff.div g.contMDiff⟩ div_eq_mul_inv := fun f g => by ext; exact div_eq_mul_inv _ _ } @[to_additive (attr := simp)] @@ -189,11 +189,11 @@ variable (I N) 'left-composition-by-`φ`' ring homomorphism from `C^∞⟮I, N; I', R'⟯` to `C^∞⟮I, N; I'', R''⟯`. -/ def compLeftRingHom {R' : Type*} [Ring R'] [TopologicalSpace R'] [ChartedSpace H' R'] [SmoothRing I' R'] {R'' : Type*} [Ring R''] [TopologicalSpace R''] [ChartedSpace H'' R''] - [SmoothRing I'' R''] (φ : R' →+* R'') (hφ : Smooth I' I'' φ) : + [SmoothRing I'' R''] (φ : R' →+* R'') (hφ : ContMDiff I' I'' ⊤ φ) : C^∞⟮I, N; I', R'⟯ →+* C^∞⟮I, N; I'', R''⟯ := { SmoothMap.compLeftMonoidHom I N φ.toMonoidHom hφ, SmoothMap.compLeftAddMonoidHom I N φ.toAddMonoidHom hφ with - toFun := fun f => ⟨φ ∘ f, fun x => (hφ.smooth _).comp x (f.contMDiff x)⟩ } + toFun := fun f => ⟨φ ∘ f, hφ.comp f.contMDiff⟩ } variable (I') {N} @@ -202,7 +202,7 @@ variable (I') {N} def restrictRingHom (R : Type*) [Ring R] [TopologicalSpace R] [ChartedSpace H' R] [SmoothRing I' R] {U V : Opens N} (h : U ≤ V) : C^∞⟮I, V; I', R⟯ →+* C^∞⟮I, U; I', R⟯ := { SmoothMap.restrictMonoidHom I I' R h, SmoothMap.restrictAddMonoidHom I I' R h with - toFun := fun f => ⟨f ∘ Set.inclusion h, f.smooth.comp (smooth_inclusion h)⟩ } + toFun := fun f => ⟨f ∘ Set.inclusion h, f.contMDiff.comp (contMDiff_inclusion h)⟩ } variable {I I'} @@ -232,7 +232,7 @@ field `𝕜` inherit a vector space structure. instance instSMul {V : Type*} [NormedAddCommGroup V] [NormedSpace 𝕜 V] : SMul 𝕜 C^∞⟮I, N; 𝓘(𝕜, V), V⟯ := - ⟨fun r f => ⟨r • ⇑f, smooth_const.smul f.smooth⟩⟩ + ⟨fun r f => ⟨r • ⇑f, contMDiff_const.smul f.contMDiff⟩⟩ @[simp] theorem coe_smul {V : Type*} [NormedAddCommGroup V] [NormedSpace 𝕜 V] (r : 𝕜) @@ -272,7 +272,7 @@ variable {A : Type*} [NormedRing A] [NormedAlgebra 𝕜 A] [SmoothRing 𝓘(𝕜 /-- Smooth constant functions as a `RingHom`. -/ def C : 𝕜 →+* C^∞⟮I, N; 𝓘(𝕜, A), A⟯ where - toFun := fun c : 𝕜 => ⟨fun _ => (algebraMap 𝕜 A) c, smooth_const⟩ + toFun := fun c : 𝕜 => ⟨fun _ => (algebraMap 𝕜 A) c, contMDiff_const⟩ map_one' := by ext; exact (algebraMap 𝕜 A).map_one map_mul' c₁ c₂ := by ext; exact (algebraMap 𝕜 A).map_mul _ _ map_zero' := by ext; exact (algebraMap 𝕜 A).map_zero @@ -280,7 +280,7 @@ def C : 𝕜 →+* C^∞⟮I, N; 𝓘(𝕜, A), A⟯ where instance algebra : Algebra 𝕜 C^∞⟮I, N; 𝓘(𝕜, A), A⟯ := { --SmoothMap.semiring with -- Porting note: Commented this out. - smul := fun r f => ⟨r • f, smooth_const.smul f.smooth⟩ + smul := fun r f => ⟨r • f, contMDiff_const.smul f.contMDiff⟩ toRingHom := SmoothMap.C commutes' := fun c f => by ext x; exact Algebra.commutes' _ _ smul_def' := fun c f => by ext x; exact Algebra.smul_def' _ _ } @@ -309,7 +309,7 @@ is naturally a vector space over the ring of smooth functions from `N` to `𝕜` instance instSMul' {V : Type*} [NormedAddCommGroup V] [NormedSpace 𝕜 V] : SMul C^∞⟮I, N; 𝕜⟯ C^∞⟮I, N; 𝓘(𝕜, V), V⟯ := - ⟨fun f g => ⟨fun x => f x • g x, Smooth.smul f.2 g.2⟩⟩ + ⟨fun f g => ⟨fun x => f x • g x, ContMDiff.smul f.2 g.2⟩⟩ @[simp] theorem smul_comp' {V : Type*} [NormedAddCommGroup V] [NormedSpace 𝕜 V] (f : C^∞⟮I'', N'; 𝕜⟯) diff --git a/Mathlib/Geometry/Manifold/Algebra/Structures.lean b/Mathlib/Geometry/Manifold/Algebra/Structures.lean index 11c9d9b1a863a..da593ec373f61 100644 --- a/Mathlib/Geometry/Manifold/Algebra/Structures.lean +++ b/Mathlib/Geometry/Manifold/Algebra/Structures.lean @@ -24,7 +24,7 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [TopologicalS If `R` is a ring, then negation is automatically smooth, as it is multiplication with `-1`. -/ class SmoothRing (I : ModelWithCorners 𝕜 E H) (R : Type*) [Semiring R] [TopologicalSpace R] [ChartedSpace H R] extends SmoothAdd I R : Prop where - smooth_mul : Smooth (I.prod I) I fun p : R × R => p.1 * p.2 + smooth_mul : ContMDiff (I.prod I) I ⊤ fun p : R × R => p.1 * p.2 -- see Note [lower instance priority] instance (priority := 100) SmoothRing.toSmoothMul (I : ModelWithCorners 𝕜 E H) (R : Type*) @@ -36,8 +36,8 @@ instance (priority := 100) SmoothRing.toSmoothMul (I : ModelWithCorners 𝕜 E H instance (priority := 100) SmoothRing.toLieAddGroup (I : ModelWithCorners 𝕜 E H) (R : Type*) [Ring R] [TopologicalSpace R] [ChartedSpace H R] [SmoothRing I R] : LieAddGroup I R where compatible := StructureGroupoid.compatible (contDiffGroupoid ⊤ I) - smooth_add := smooth_add I - smooth_neg := by simpa only [neg_one_mul] using @smooth_mul_left 𝕜 _ H _ E _ _ I R _ _ _ _ (-1) + smooth_add := contMDiff_add I + smooth_neg := by simpa only [neg_one_mul] using contMDiff_mul_left (G := R) (a := -1) end SmoothRing @@ -46,7 +46,7 @@ instance (priority := 100) fieldSmoothRing {𝕜 : Type*} [NontriviallyNormedFie SmoothRing 𝓘(𝕜) 𝕜 := { normedSpaceLieAddGroup with smooth_mul := by - rw [smooth_iff] + rw [contMDiff_iff] refine ⟨continuous_mul, fun x y => ?_⟩ simp only [mfld_simps] rw [contDiffOn_univ] diff --git a/Mathlib/Geometry/Manifold/BumpFunction.lean b/Mathlib/Geometry/Manifold/BumpFunction.lean index fa6698b53773a..0c37faf328e8a 100644 --- a/Mathlib/Geometry/Manifold/BumpFunction.lean +++ b/Mathlib/Geometry/Manifold/BumpFunction.lean @@ -286,23 +286,28 @@ theorem nhds_basis_support {s : Set M} (hs : s ∈ 𝓝 c) : variable [SmoothManifoldWithCorners I M] /-- A smooth bump function is infinitely smooth. -/ -protected theorem smooth : Smooth I 𝓘(ℝ) f := by +protected theorem contMDiff : ContMDiff I 𝓘(ℝ) ⊤ f := by refine contMDiff_of_tsupport fun x hx => ?_ have : x ∈ (chartAt H c).source := f.tsupport_subset_chartAt_source hx refine ContMDiffAt.congr_of_eventuallyEq ?_ <| f.eqOn_source.eventuallyEq_of_mem <| (chartAt H c).open_source.mem_nhds this exact f.contDiffAt.contMDiffAt.comp _ (contMDiffAt_extChartAt' this) -protected theorem smoothAt {x} : SmoothAt I 𝓘(ℝ) f x := - f.smooth.smoothAt +@[deprecated (since := "2024-11-20")] alias smooth := SmoothBumpFunction.contMDiff + +protected theorem contMDiffAt {x} : ContMDiffAt I 𝓘(ℝ) ⊤ f x := + f.contMDiff.contMDiffAt + +@[deprecated (since := "2024-11-20")] alias smoothAt := SmoothBumpFunction.contMDiffAt protected theorem continuous : Continuous f := - f.smooth.continuous + f.contMDiff.continuous /-- If `f : SmoothBumpFunction I c` is a smooth bump function and `g : M → G` is a function smooth on the source of the chart at `c`, then `f • g` is smooth on the whole manifold. -/ -theorem smooth_smul {G} [NormedAddCommGroup G] [NormedSpace ℝ G] {g : M → G} - (hg : SmoothOn I 𝓘(ℝ, G) g (chartAt H c).source) : Smooth I 𝓘(ℝ, G) fun x => f x • g x := by +theorem contMDiff_smul {G} [NormedAddCommGroup G] [NormedSpace ℝ G] {g : M → G} + (hg : ContMDiffOn I 𝓘(ℝ, G) ⊤ g (chartAt H c).source) : + ContMDiff I 𝓘(ℝ, G) ⊤ fun x => f x • g x := by refine contMDiff_of_tsupport fun x hx => ?_ have : x ∈ (chartAt H c).source := -- Porting note: was a more readable `calc` @@ -311,6 +316,8 @@ theorem smooth_smul {G} [NormedAddCommGroup G] [NormedSpace ℝ G] {g : M → G} -- _ ⊆ tsupport f := tsupport_smul_subset_left _ _ -- _ ⊆ (chart_at _ c).source := f.tsupport_subset_chartAt_source f.tsupport_subset_chartAt_source <| tsupport_smul_subset_left _ _ hx - exact f.smoothAt.smul ((hg _ this).contMDiffAt <| (chartAt _ _).open_source.mem_nhds this) + exact f.contMDiffAt.smul ((hg _ this).contMDiffAt <| (chartAt _ _).open_source.mem_nhds this) + +@[deprecated (since := "2024-11-20")] alias smooth_smul := contMDiff_smul end SmoothBumpFunction diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean b/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean index 5f68ad7682a06..2036cda540d21 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean @@ -125,7 +125,7 @@ variable [ChartedSpace H M'] [IsM' : SmoothManifoldWithCorners I M'] theorem isLocalStructomorphOn_contDiffGroupoid_iff_aux {f : PartialHomeomorph M M'} (hf : LiftPropOn (contDiffGroupoid ⊤ I).IsLocalStructomorphWithinAt f f.source) : - SmoothOn I I f f.source := by + ContMDiffOn I I ⊤ f f.source := by -- It suffices to show smoothness near each `x` apply contMDiffOn_of_locally_contMDiffOn intro x hx @@ -171,7 +171,7 @@ is a local structomorphism for `I`, if and only if it is manifold-smooth on the in both directions. -/ theorem isLocalStructomorphOn_contDiffGroupoid_iff (f : PartialHomeomorph M M') : LiftPropOn (contDiffGroupoid ⊤ I).IsLocalStructomorphWithinAt f f.source ↔ - SmoothOn I I f f.source ∧ SmoothOn I I f.symm f.target := by + ContMDiffOn I I ⊤ f f.source ∧ ContMDiffOn I I ⊤ f.symm f.target := by constructor · intro h refine ⟨isLocalStructomorphOn_contDiffGroupoid_iff_aux h, diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean b/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean index 0b2a089aee1e6..f5506bc51360d 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean @@ -82,31 +82,21 @@ theorem ContMDiffWithinAt.comp_of_eq {t : Set M'} {g : M' → M''} {x : M} {y : (st : MapsTo f s t) (hx : f x = y) : ContMDiffWithinAt I I'' n (g ∘ f) s x := by subst hx; exact hg.comp x hf st -/-- The composition of `C^∞` functions within domains at points is `C^∞`. -/ -nonrec theorem SmoothWithinAt.comp {t : Set M'} {g : M' → M''} (x : M) - (hg : SmoothWithinAt I' I'' g t (f x)) (hf : SmoothWithinAt I I' f s x) (st : MapsTo f s t) : - SmoothWithinAt I I'' (g ∘ f) s x := - hg.comp x hf st +@[deprecated (since := "2024-11-20")] alias SmoothWithinAt.comp := ContMDiffWithinAt.comp /-- The composition of `C^n` functions on domains is `C^n`. -/ theorem ContMDiffOn.comp {t : Set M'} {g : M' → M''} (hg : ContMDiffOn I' I'' n g t) (hf : ContMDiffOn I I' n f s) (st : s ⊆ f ⁻¹' t) : ContMDiffOn I I'' n (g ∘ f) s := fun x hx => (hg _ (st hx)).comp x (hf x hx) st -/-- The composition of `C^∞` functions on domains is `C^∞`. -/ -nonrec theorem SmoothOn.comp {t : Set M'} {g : M' → M''} (hg : SmoothOn I' I'' g t) - (hf : SmoothOn I I' f s) (st : s ⊆ f ⁻¹' t) : SmoothOn I I'' (g ∘ f) s := - hg.comp hf st +@[deprecated (since := "2024-11-20")] alias SmoothOn.comp := ContMDiffOn.comp /-- The composition of `C^n` functions on domains is `C^n`. -/ theorem ContMDiffOn.comp' {t : Set M'} {g : M' → M''} (hg : ContMDiffOn I' I'' n g t) (hf : ContMDiffOn I I' n f s) : ContMDiffOn I I'' n (g ∘ f) (s ∩ f ⁻¹' t) := hg.comp (hf.mono inter_subset_left) inter_subset_right -/-- The composition of `C^∞` functions is `C^∞`. -/ -nonrec theorem SmoothOn.comp' {t : Set M'} {g : M' → M''} (hg : SmoothOn I' I'' g t) - (hf : SmoothOn I I' f s) : SmoothOn I I'' (g ∘ f) (s ∩ f ⁻¹' t) := - hg.comp' hf +@[deprecated (since := "2024-11-20")] alias SmoothOn.comp' := ContMDiffOn.comp' /-- The composition of `C^n` functions is `C^n`. -/ theorem ContMDiff.comp {g : M' → M''} (hg : ContMDiff I' I'' n g) (hf : ContMDiff I I' n f) : @@ -114,10 +104,7 @@ theorem ContMDiff.comp {g : M' → M''} (hg : ContMDiff I' I'' n g) (hf : ContMD rw [← contMDiffOn_univ] at hf hg ⊢ exact hg.comp hf subset_preimage_univ -/-- The composition of `C^∞` functions is `C^∞`. -/ -nonrec theorem Smooth.comp {g : M' → M''} (hg : Smooth I' I'' g) (hf : Smooth I I' f) : - Smooth I I'' (g ∘ f) := - hg.comp hf +@[deprecated (since := "2024-11-20")] alias Smooth.comp := ContMDiff.comp /-- The composition of `C^n` functions within domains at points is `C^n`. -/ theorem ContMDiffWithinAt.comp' {t : Set M'} {g : M' → M''} (x : M) @@ -125,11 +112,7 @@ theorem ContMDiffWithinAt.comp' {t : Set M'} {g : M' → M''} (x : M) ContMDiffWithinAt I I'' n (g ∘ f) (s ∩ f ⁻¹' t) x := hg.comp x (hf.mono inter_subset_left) inter_subset_right -/-- The composition of `C^∞` functions within domains at points is `C^∞`. -/ -nonrec theorem SmoothWithinAt.comp' {t : Set M'} {g : M' → M''} (x : M) - (hg : SmoothWithinAt I' I'' g t (f x)) (hf : SmoothWithinAt I I' f s x) : - SmoothWithinAt I I'' (g ∘ f) (s ∩ f ⁻¹' t) x := - hg.comp' x hf +@[deprecated (since := "2024-11-20")] alias SmoothWithinAt.comp' := ContMDiffWithinAt.comp' /-- `g ∘ f` is `C^n` within `s` at `x` if `g` is `C^n` at `f x` and `f` is `C^n` within `s` at `x`. -/ @@ -138,11 +121,8 @@ theorem ContMDiffAt.comp_contMDiffWithinAt {g : M' → M''} (x : M) ContMDiffWithinAt I I'' n (g ∘ f) s x := hg.comp x hf (mapsTo_univ _ _) -/-- `g ∘ f` is `C^∞` within `s` at `x` if `g` is `C^∞` at `f x` and -`f` is `C^∞` within `s` at `x`. -/ -theorem SmoothAt.comp_smoothWithinAt {g : M' → M''} (x : M) (hg : SmoothAt I' I'' g (f x)) - (hf : SmoothWithinAt I I' f s x) : SmoothWithinAt I I'' (g ∘ f) s x := - hg.comp_contMDiffWithinAt x hf +@[deprecated (since := "2024-11-20")] +alias SmoothAt.comp_smoothWithinAt := ContMDiffAt.comp_contMDiffWithinAt /-- The composition of `C^n` functions at points is `C^n`. -/ nonrec theorem ContMDiffAt.comp {g : M' → M''} (x : M) (hg : ContMDiffAt I' I'' n g (f x)) @@ -154,26 +134,19 @@ theorem ContMDiffAt.comp_of_eq {g : M' → M''} {x : M} {y : M'} (hg : ContMDiff (hf : ContMDiffAt I I' n f x) (hx : f x = y) : ContMDiffAt I I'' n (g ∘ f) x := by subst hx; exact hg.comp x hf -/-- The composition of `C^∞` functions at points is `C^∞`. -/ -nonrec theorem SmoothAt.comp {g : M' → M''} (x : M) (hg : SmoothAt I' I'' g (f x)) - (hf : SmoothAt I I' f x) : SmoothAt I I'' (g ∘ f) x := - hg.comp x hf +@[deprecated (since := "2024-11-20")] alias SmoothAt.comp := ContMDiffAt.comp theorem ContMDiff.comp_contMDiffOn {f : M → M'} {g : M' → M''} {s : Set M} (hg : ContMDiff I' I'' n g) (hf : ContMDiffOn I I' n f s) : ContMDiffOn I I'' n (g ∘ f) s := hg.contMDiffOn.comp hf Set.subset_preimage_univ -theorem Smooth.comp_smoothOn {f : M → M'} {g : M' → M''} {s : Set M} (hg : Smooth I' I'' g) - (hf : SmoothOn I I' f s) : SmoothOn I I'' (g ∘ f) s := - hg.smoothOn.comp hf Set.subset_preimage_univ +@[deprecated (since := "2024-11-20")] alias Smooth.comp_smoothOn := ContMDiff.comp_contMDiffOn theorem ContMDiffOn.comp_contMDiff {t : Set M'} {g : M' → M''} (hg : ContMDiffOn I' I'' n g t) (hf : ContMDiff I I' n f) (ht : ∀ x, f x ∈ t) : ContMDiff I I'' n (g ∘ f) := contMDiffOn_univ.mp <| hg.comp hf.contMDiffOn fun x _ => ht x -theorem SmoothOn.comp_smooth {t : Set M'} {g : M' → M''} (hg : SmoothOn I' I'' g t) - (hf : Smooth I I' f) (ht : ∀ x, f x ∈ t) : Smooth I I'' (g ∘ f) := - hg.comp_contMDiff hf ht +@[deprecated (since := "2024-11-20")] alias SmoothOn.comp_smooth := ContMDiffOn.comp_contMDiff end Composition @@ -185,26 +158,22 @@ theorem contMDiff_id : ContMDiff I I n (id : M → M) := ContMDiff.of_le ((contDiffWithinAt_localInvariantProp ∞).liftProp_id contDiffWithinAtProp_id) le_top -theorem smooth_id : Smooth I I (id : M → M) := - contMDiff_id +@[deprecated (since := "2024-11-20")] alias smooth_id := contMDiff_id theorem contMDiffOn_id : ContMDiffOn I I n (id : M → M) s := contMDiff_id.contMDiffOn -theorem smoothOn_id : SmoothOn I I (id : M → M) s := - contMDiffOn_id +@[deprecated (since := "2024-11-20")] alias smoothOn_id := contMDiffOn_id theorem contMDiffAt_id : ContMDiffAt I I n (id : M → M) x := contMDiff_id.contMDiffAt -theorem smoothAt_id : SmoothAt I I (id : M → M) x := - contMDiffAt_id +@[deprecated (since := "2024-11-20")] alias smoothAt_id := contMDiffAt_id theorem contMDiffWithinAt_id : ContMDiffWithinAt I I n (id : M → M) s x := contMDiffAt_id.contMDiffWithinAt -theorem smoothWithinAt_id : SmoothWithinAt I I (id : M → M) s x := - contMDiffWithinAt_id +@[deprecated (since := "2024-11-20")] alias smoothWithinAt_id := contMDiffWithinAt_id end id @@ -223,11 +192,10 @@ theorem contMDiff_const : ContMDiff I I' n fun _ : M => c := by theorem contMDiff_one [One M'] : ContMDiff I I' n (1 : M → M') := by simp only [Pi.one_def, contMDiff_const] -theorem smooth_const : Smooth I I' fun _ : M => c := - contMDiff_const +@[deprecated (since := "2024-11-20")] alias smooth_const := contMDiff_const -@[to_additive] -theorem smooth_one [One M'] : Smooth I I' (1 : M → M') := by simp only [Pi.one_def, smooth_const] +@[deprecated (since := "2024-11-20")] alias smooth_one := contMDiff_one +@[deprecated (since := "2024-11-20")] alias smooth_zero := contMDiff_zero theorem contMDiffOn_const : ContMDiffOn I I' n (fun _ : M => c) s := contMDiff_const.contMDiffOn @@ -236,12 +204,11 @@ theorem contMDiffOn_const : ContMDiffOn I I' n (fun _ : M => c) s := theorem contMDiffOn_one [One M'] : ContMDiffOn I I' n (1 : M → M') s := contMDiff_one.contMDiffOn -theorem smoothOn_const : SmoothOn I I' (fun _ : M => c) s := - contMDiffOn_const +@[deprecated (since := "2024-11-20")] alias smoothOn_const := contMDiffOn_const + +@[deprecated (since := "2024-11-20")] alias smoothOn_one := contMDiffOn_one +@[deprecated (since := "2024-11-20")] alias smoothOn_zero := contMDiffOn_zero -@[to_additive] -theorem smoothOn_one [One M'] : SmoothOn I I' (1 : M → M') s := - contMDiffOn_one theorem contMDiffAt_const : ContMDiffAt I I' n (fun _ : M => c) x := contMDiff_const.contMDiffAt @@ -250,12 +217,10 @@ theorem contMDiffAt_const : ContMDiffAt I I' n (fun _ : M => c) x := theorem contMDiffAt_one [One M'] : ContMDiffAt I I' n (1 : M → M') x := contMDiff_one.contMDiffAt -theorem smoothAt_const : SmoothAt I I' (fun _ : M => c) x := - contMDiffAt_const +@[deprecated (since := "2024-11-20")] alias smoothAt_const := contMDiffAt_const -@[to_additive] -theorem smoothAt_one [One M'] : SmoothAt I I' (1 : M → M') x := - contMDiffAt_one +@[deprecated (since := "2024-11-20")] alias smoothAt_one := contMDiffAt_one +@[deprecated (since := "2024-11-20")] alias smoothAt_zero := contMDiffAt_zero theorem contMDiffWithinAt_const : ContMDiffWithinAt I I' n (fun _ : M => c) s x := contMDiffAt_const.contMDiffWithinAt @@ -264,12 +229,10 @@ theorem contMDiffWithinAt_const : ContMDiffWithinAt I I' n (fun _ : M => c) s x theorem contMDiffWithinAt_one [One M'] : ContMDiffWithinAt I I' n (1 : M → M') s x := contMDiffAt_const.contMDiffWithinAt -theorem smoothWithinAt_const : SmoothWithinAt I I' (fun _ : M => c) s x := - contMDiffWithinAt_const +@[deprecated (since := "2024-11-20")] alias smoothWithinAt_const := contMDiffWithinAt_const -@[to_additive] -theorem smoothWithinAt_one [One M'] : SmoothWithinAt I I' (1 : M → M') s x := - contMDiffWithinAt_one +@[deprecated (since := "2024-11-20")] alias smoothWithinAt_one := contMDiffWithinAt_one +@[deprecated (since := "2024-11-20")] alias smoothWithinAt_zero := contMDiffWithinAt_zero end id @@ -305,12 +268,14 @@ section Inclusion open TopologicalSpace -theorem contMdiffAt_subtype_iff {n : ℕ∞} {U : Opens M} {f : M → M'} {x : U} : +theorem contMDiffAt_subtype_iff {n : ℕ∞} {U : Opens M} {f : M → M'} {x : U} : ContMDiffAt I I' n (fun x : U ↦ f x) x ↔ ContMDiffAt I I' n f x := ((contDiffWithinAt_localInvariantProp n).liftPropAt_iff_comp_subtype_val _ _).symm +@[deprecated (since := "2024-11-20")] alias contMdiffAt_subtype_iff := contMDiffAt_subtype_iff + theorem contMDiff_subtype_val {n : ℕ∞} {U : Opens M} : ContMDiff I I n (Subtype.val : U → M) := - fun _ ↦ contMdiffAt_subtype_iff.mpr contMDiffAt_id + fun _ ↦ contMDiffAt_subtype_iff.mpr contMDiffAt_id @[to_additive] theorem ContMDiff.extend_one [T2Space M] [One M'] {n : ℕ∞} {U : Opens M} {f : U → M'} @@ -319,7 +284,7 @@ theorem ContMDiff.extend_one [T2Space M] [One M'] {n : ℕ∞} {U : Opens M} {f refine contMDiff_of_mulTSupport (fun x h ↦ ?_) _ lift x to U using Subtype.coe_image_subset _ _ (supp.mulTSupport_extend_one_subset continuous_subtype_val h) - rw [← contMdiffAt_subtype_iff] + rw [← contMDiffAt_subtype_iff] simp_rw [← comp_def] rw [extend_comp Subtype.val_injective] exact diff.contMDiffAt @@ -333,19 +298,14 @@ theorem contMDiff_inclusion {n : ℕ∞} {U V : Opens M} (h : U ≤ V) : rw [Set.univ_inter] exact contDiffWithinAt_id.congr I.rightInvOn (congr_arg I (I.left_inv y)) -theorem smooth_subtype_iff {U : Opens M} {f : M → M'} {x : U} : - SmoothAt I I' (fun x : U ↦ f x) x ↔ SmoothAt I I' f x := contMdiffAt_subtype_iff +@[deprecated (since := "2024-11-20")] alias smooth_subtype_iff := contMDiffAt_subtype_iff -theorem smooth_subtype_val {U : Opens M} : Smooth I I (Subtype.val : U → M) := contMDiff_subtype_val +@[deprecated (since := "2024-11-20")] alias smooth_subtype_val := contMDiff_subtype_val -@[to_additive] -theorem Smooth.extend_one [T2Space M] [One M'] {U : Opens M} {f : U → M'} - (supp : HasCompactMulSupport f) (diff : Smooth I I' f) : Smooth I I' (Subtype.val.extend f 1) := - ContMDiff.extend_one supp diff +@[deprecated (since := "2024-11-20")] alias Smooth.extend_one := ContMDiff.extend_one +@[deprecated (since := "2024-11-20")] alias Smooth.extend_zero := ContMDiff.extend_zero -theorem smooth_inclusion {U V : Opens M} (h : U ≤ V) : - Smooth I I (Opens.inclusion h : U → V) := - contMDiff_inclusion h +@[deprecated (since := "2024-11-20")] alias smooth_inclusion := contMDiff_inclusion end Inclusion diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean index f0917e558dc3c..aa5e8d957cfd7 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean @@ -155,11 +155,7 @@ read in the preferred chart at this point. -/ def ContMDiffWithinAt (n : ℕ∞) (f : M → M') (s : Set M) (x : M) := LiftPropWithinAt (ContDiffWithinAtProp I I' n) f s x -variable (I I') in -/-- Abbreviation for `ContMDiffWithinAt I I' ⊤ f s x`. See also documentation for `Smooth`. --/ -abbrev SmoothWithinAt (f : M → M') (s : Set M) (x : M) := - ContMDiffWithinAt I I' ⊤ f s x +@[deprecated (since := "024-11-21")] alias SmoothWithinAt := ContMDiffWithinAt variable (I I') in /-- A function is `n` times continuously differentiable at a point in a manifold if @@ -175,10 +171,7 @@ theorem contMDiffAt_iff {n : ℕ∞} {f : M → M'} {x : M} : (extChartAt I x x) := liftPropAt_iff.trans <| by rw [ContDiffWithinAtProp, preimage_univ, univ_inter]; rfl -variable (I I') in -/-- Abbreviation for `ContMDiffAt I I' ⊤ f x`. See also documentation for `Smooth`. -/ -abbrev SmoothAt (f : M → M') (x : M) := - ContMDiffAt I I' ⊤ f x +@[deprecated (since := "024-11-21")] alias SmoothAt := ContMDiffAt variable (I I') in /-- A function is `n` times continuously differentiable in a set of a manifold if it is continuous @@ -187,10 +180,7 @@ around these points. -/ def ContMDiffOn (n : ℕ∞) (f : M → M') (s : Set M) := ∀ x ∈ s, ContMDiffWithinAt I I' n f s x -variable (I I') in -/-- Abbreviation for `ContMDiffOn I I' ⊤ f s`. See also documentation for `Smooth`. -/ -abbrev SmoothOn (f : M → M') (s : Set M) := - ContMDiffOn I I' ⊤ f s +@[deprecated (since := "024-11-21")] alias SmoothOn := ContMDiffOn variable (I I') in /-- A function is `n` times continuously differentiable in a manifold if it is continuous @@ -199,16 +189,8 @@ around these points. -/ def ContMDiff (n : ℕ∞) (f : M → M') := ∀ x, ContMDiffAt I I' n f x -variable (I I') in -/-- Abbreviation for `ContMDiff I I' ⊤ f`. -Short note to work with these abbreviations: a lemma of the form `ContMDiffFoo.bar` will -apply fine to an assumption `SmoothFoo` using dot notation or normal notation. -If the consequence `bar` of the lemma involves `ContDiff`, it is still better to restate -the lemma replacing `ContDiff` with `Smooth` both in the assumption and in the conclusion, -to make it possible to use `Smooth` consistently. -This also applies to `SmoothAt`, `SmoothOn` and `SmoothWithinAt`. -/ -abbrev Smooth (f : M → M') := - ContMDiff I I' ⊤ f +@[deprecated (since := "024-11-21")] alias Smooth := ContMDiff + /-! ### Deducing smoothness from higher smoothness -/ @@ -228,49 +210,38 @@ theorem ContMDiff.of_le (hf : ContMDiff I I' n f) (le : m ≤ n) : ContMDiff I I /-! ### Basic properties of smooth functions between manifolds -/ -theorem ContMDiff.smooth (h : ContMDiff I I' ⊤ f) : Smooth I I' f := - h +@[deprecated (since := "2024-11-20")] alias ContMDiff.smooth := ContMDiff.of_le -theorem Smooth.contMDiff (h : Smooth I I' f) : ContMDiff I I' n f := - h.of_le le_top +@[deprecated (since := "2024-11-20")] alias Smooth.contMDiff := ContMDiff.of_le -theorem ContMDiffOn.smoothOn (h : ContMDiffOn I I' ⊤ f s) : SmoothOn I I' f s := - h +@[deprecated (since := "2024-11-20")] alias ContMDiffOn.smoothOn := ContMDiffOn.of_le -theorem SmoothOn.contMDiffOn (h : SmoothOn I I' f s) : ContMDiffOn I I' n f s := - h.of_le le_top +@[deprecated (since := "2024-11-20")] alias SmoothOn.contMDiffOn := ContMDiffOn.of_le -theorem ContMDiffAt.smoothAt (h : ContMDiffAt I I' ⊤ f x) : SmoothAt I I' f x := - h +@[deprecated (since := "2024-11-20")] alias ContMDiffAt.smoothAt := ContMDiffAt.of_le -theorem SmoothAt.contMDiffAt (h : SmoothAt I I' f x) : ContMDiffAt I I' n f x := - h.of_le le_top +@[deprecated (since := "2024-11-20")] alias SmoothAt.contMDiffAt := ContMDiffOn.of_le -theorem ContMDiffWithinAt.smoothWithinAt (h : ContMDiffWithinAt I I' ⊤ f s x) : - SmoothWithinAt I I' f s x := - h +@[deprecated (since := "2024-11-20")] +alias ContMDiffWithinAt.smoothWithinAt := ContMDiffWithinAt.of_le -theorem SmoothWithinAt.contMDiffWithinAt (h : SmoothWithinAt I I' f s x) : - ContMDiffWithinAt I I' n f s x := - h.of_le le_top +@[deprecated (since := "2024-11-20")] +alias SmoothWithinAt.contMDiffWithinAt := ContMDiffWithinAt.of_le theorem ContMDiff.contMDiffAt (h : ContMDiff I I' n f) : ContMDiffAt I I' n f x := h x -theorem Smooth.smoothAt (h : Smooth I I' f) : SmoothAt I I' f x := - ContMDiff.contMDiffAt h +@[deprecated (since := "2024-11-20")] alias Smooth.smoothAt := ContMDiff.contMDiffAt theorem contMDiffWithinAt_univ : ContMDiffWithinAt I I' n f univ x ↔ ContMDiffAt I I' n f x := Iff.rfl -theorem smoothWithinAt_univ : SmoothWithinAt I I' f univ x ↔ SmoothAt I I' f x := - contMDiffWithinAt_univ +@[deprecated (since := "2024-11-20")] alias smoothWithinAt_univ := contMDiffWithinAt_univ theorem contMDiffOn_univ : ContMDiffOn I I' n f univ ↔ ContMDiff I I' n f := by simp only [ContMDiffOn, ContMDiff, contMDiffWithinAt_univ, forall_prop_of_true, mem_univ] -theorem smoothOn_univ : SmoothOn I I' f univ ↔ Smooth I I' f := - contMDiffOn_univ +@[deprecated (since := "2024-11-20")] alias smoothOn_univ := contMDiffOn_univ /-- One can reformulate smoothness within a set at a point as continuity within this set at this point, and smoothness in the corresponding extended chart. -/ @@ -315,26 +286,18 @@ theorem contMDiffWithinAt_iff_target : chartAt_self_eq, PartialHomeomorph.refl_apply, id_comp] rfl -theorem smoothWithinAt_iff : - SmoothWithinAt I I' f s x ↔ - ContinuousWithinAt f s x ∧ - ContDiffWithinAt 𝕜 ∞ (extChartAt I' (f x) ∘ f ∘ (extChartAt I x).symm) - ((extChartAt I x).symm ⁻¹' s ∩ range I) (extChartAt I x x) := - contMDiffWithinAt_iff +@[deprecated (since := "2024-11-20")] alias smoothWithinAt_iff := contMDiffWithinAt_iff -theorem smoothWithinAt_iff_target : - SmoothWithinAt I I' f s x ↔ - ContinuousWithinAt f s x ∧ SmoothWithinAt I 𝓘(𝕜, E') (extChartAt I' (f x) ∘ f) s x := - contMDiffWithinAt_iff_target +@[deprecated (since := "2024-11-20")] +alias smoothWithinAt_iff_target := contMDiffWithinAt_iff_target theorem contMDiffAt_iff_target {x : M} : ContMDiffAt I I' n f x ↔ ContinuousAt f x ∧ ContMDiffAt I 𝓘(𝕜, E') n (extChartAt I' (f x) ∘ f) x := by rw [ContMDiffAt, ContMDiffAt, contMDiffWithinAt_iff_target, continuousWithinAt_univ] -theorem smoothAt_iff_target {x : M} : - SmoothAt I I' f x ↔ ContinuousAt f x ∧ SmoothAt I 𝓘(𝕜, E') (extChartAt I' (f x) ∘ f) x := - contMDiffAt_iff_target +@[deprecated (since := "2024-11-20")] alias smoothAt_iff_target := contMDiffAt_iff_target + section SmoothManifoldWithCorners @@ -533,20 +496,10 @@ theorem contMDiffOn_iff_target : simp · exact fun h' x y => (h' y).2 x 0 -theorem smoothOn_iff : - SmoothOn I I' f s ↔ - ContinuousOn f s ∧ - ∀ (x : M) (y : M'), - ContDiffOn 𝕜 ⊤ (extChartAt I' y ∘ f ∘ (extChartAt I x).symm) - ((extChartAt I x).target ∩ - (extChartAt I x).symm ⁻¹' (s ∩ f ⁻¹' (extChartAt I' y).source)) := - contMDiffOn_iff +@[deprecated (since := "2024-11-20")] alias smoothOn_iff := contMDiffOn_iff + +@[deprecated (since := "2024-11-20")] alias smoothOn_iff_target := contMDiffOn_iff_target -theorem smoothOn_iff_target : - SmoothOn I I' f s ↔ - ContinuousOn f s ∧ - ∀ y : M', SmoothOn I 𝓘(𝕜, E') (extChartAt I' y ∘ f) (s ∩ f ⁻¹' (extChartAt I' y).source) := - contMDiffOn_iff_target /-- One can reformulate smoothness as continuity and smoothness in any extended chart. -/ theorem contMDiff_iff : @@ -567,20 +520,9 @@ theorem contMDiff_iff_target : rw [← contMDiffOn_univ, contMDiffOn_iff_target] simp [continuous_iff_continuousOn_univ] -theorem smooth_iff : - Smooth I I' f ↔ - Continuous f ∧ - ∀ (x : M) (y : M'), - ContDiffOn 𝕜 ⊤ (extChartAt I' y ∘ f ∘ (extChartAt I x).symm) - ((extChartAt I x).target ∩ - (extChartAt I x).symm ⁻¹' (f ⁻¹' (extChartAt I' y).source)) := - contMDiff_iff +@[deprecated (since := "2024-11-20")] alias smooth_iff := contMDiff_iff -theorem smooth_iff_target : - Smooth I I' f ↔ - Continuous f ∧ - ∀ y : M', SmoothOn I 𝓘(𝕜, E') (extChartAt I' y ∘ f) (f ⁻¹' (extChartAt I' y).source) := - contMDiff_iff_target +@[deprecated (since := "2024-11-20")] alias smooth_iff_target := contMDiff_iff_target end SmoothManifoldWithCorners @@ -619,17 +561,17 @@ theorem ContMDiff.continuous (hf : ContMDiff I I' n f) : Continuous f := /-! ### `C^∞` smoothness -/ theorem contMDiffWithinAt_top : - SmoothWithinAt I I' f s x ↔ ∀ n : ℕ, ContMDiffWithinAt I I' n f s x := + ContMDiffWithinAt I I' ⊤ f s x ↔ ∀ n : ℕ, ContMDiffWithinAt I I' n f s x := ⟨fun h n => ⟨h.1, contDiffWithinAt_top.1 h.2 n⟩, fun H => ⟨(H 0).1, contDiffWithinAt_top.2 fun n => (H n).2⟩⟩ -theorem contMDiffAt_top : SmoothAt I I' f x ↔ ∀ n : ℕ, ContMDiffAt I I' n f x := +theorem contMDiffAt_top : ContMDiffAt I I' ⊤ f x ↔ ∀ n : ℕ, ContMDiffAt I I' n f x := contMDiffWithinAt_top -theorem contMDiffOn_top : SmoothOn I I' f s ↔ ∀ n : ℕ, ContMDiffOn I I' n f s := +theorem contMDiffOn_top : ContMDiffOn I I' ⊤ f s ↔ ∀ n : ℕ, ContMDiffOn I I' n f s := ⟨fun h _ => h.of_le le_top, fun h x hx => contMDiffWithinAt_top.2 fun n => h n x hx⟩ -theorem contMDiff_top : Smooth I I' f ↔ ∀ n : ℕ, ContMDiff I I' n f := +theorem contMDiff_top : ContMDiff I I' ⊤ f ↔ ∀ n : ℕ, ContMDiff I I' n f := ⟨fun h _ => h.of_le le_top, fun h x => contMDiffWithinAt_top.2 fun n => h n x⟩ theorem contMDiffWithinAt_iff_nat : @@ -691,8 +633,7 @@ protected theorem ContMDiffAt.contMDiffWithinAt (hf : ContMDiffAt I I' n f x) : ContMDiffWithinAt I I' n f s x := ContMDiffWithinAt.mono hf (subset_univ _) -protected theorem SmoothAt.smoothWithinAt (hf : SmoothAt I I' f x) : SmoothWithinAt I I' f s x := - ContMDiffAt.contMDiffWithinAt hf +@[deprecated (since := "2024-11-20")] alias SmoothAt.smoothWithinAt := ContMDiffAt.contMDiffWithinAt theorem ContMDiffOn.mono (hf : ContMDiffOn I I' n f s) (hts : t ⊆ s) : ContMDiffOn I I' n f t := fun x hx => (hf x (hts hx)).mono hts @@ -700,8 +641,7 @@ theorem ContMDiffOn.mono (hf : ContMDiffOn I I' n f s) (hts : t ⊆ s) : ContMDi protected theorem ContMDiff.contMDiffOn (hf : ContMDiff I I' n f) : ContMDiffOn I I' n f s := fun x _ => (hf x).contMDiffWithinAt -protected theorem Smooth.smoothOn (hf : Smooth I I' f) : SmoothOn I I' f s := - ContMDiff.contMDiffOn hf +@[deprecated (since := "2024-11-20")] alias Smooth.smoothOn := ContMDiff.contMDiffOn theorem contMDiffWithinAt_inter' (ht : t ∈ 𝓝[s] x) : ContMDiffWithinAt I I' n f (s ∩ t) x ↔ ContMDiffWithinAt I I' n f s x := @@ -716,16 +656,13 @@ protected theorem ContMDiffWithinAt.contMDiffAt ContMDiffAt I I' n f x := (contDiffWithinAt_localInvariantProp n).liftPropAt_of_liftPropWithinAt h ht -protected theorem SmoothWithinAt.smoothAt (h : SmoothWithinAt I I' f s x) (ht : s ∈ 𝓝 x) : - SmoothAt I I' f x := - ContMDiffWithinAt.contMDiffAt h ht +@[deprecated (since := "2024-11-20")] alias SmoothWithinAt.smoothAt := ContMDiffWithinAt.contMDiffAt protected theorem ContMDiffOn.contMDiffAt (h : ContMDiffOn I I' n f s) (hx : s ∈ 𝓝 x) : ContMDiffAt I I' n f x := (h x (mem_of_mem_nhds hx)).contMDiffAt hx -protected theorem SmoothOn.smoothAt (h : SmoothOn I I' f s) (hx : s ∈ 𝓝 x) : SmoothAt I I' f x := - h.contMDiffAt hx +@[deprecated (since := "2024-11-20")] alias SmoothOn.smoothAt := ContMDiffOn.contMDiffAt theorem contMDiffOn_iff_source_of_mem_maximalAtlas [SmoothManifoldWithCorners I M] (he : e ∈ maximalAtlas I M) (hs : s ⊆ e.source) : diff --git a/Mathlib/Geometry/Manifold/ContMDiff/NormedSpace.lean b/Mathlib/Geometry/Manifold/ContMDiff/NormedSpace.lean index de822fbecd69c..e11bb6f43fe1c 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/NormedSpace.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/NormedSpace.lean @@ -105,7 +105,8 @@ theorem ContinuousLinearMap.contMDiffWithinAt (L : E →L[𝕜] F) {s x} : theorem ContinuousLinearMap.contMDiffOn (L : E →L[𝕜] F) {s} : ContMDiffOn 𝓘(𝕜, E) 𝓘(𝕜, F) n L s := L.contMDiff.contMDiffOn -theorem ContinuousLinearMap.smooth (L : E →L[𝕜] F) : Smooth 𝓘(𝕜, E) 𝓘(𝕜, F) L := L.contMDiff +@[deprecated (since := "2024-11-20")] +alias ContinuousLinearMap.smooth := ContinuousLinearMap.contMDiff theorem ContMDiffWithinAt.clm_precomp {f : M → F₁ →L[𝕜] F₂} {s : Set M} {x : M} (hf : ContMDiffWithinAt I 𝓘(𝕜, F₁ →L[𝕜] F₂) n f s x) : @@ -261,13 +262,15 @@ theorem ContMDiff.clm_prodMap {g : M → F₁ →L[𝕜] F₃} {f : M → F₂ variable {V : Type*} [NormedAddCommGroup V] [NormedSpace 𝕜 V] /-- On any vector space, multiplication by a scalar is a smooth operation. -/ -theorem smooth_smul : Smooth (𝓘(𝕜).prod 𝓘(𝕜, V)) 𝓘(𝕜, V) fun p : 𝕜 × V => p.1 • p.2 := - smooth_iff.2 ⟨continuous_smul, fun _ _ => contDiff_smul.contDiffOn⟩ +theorem contMDiff_smul : ContMDiff (𝓘(𝕜).prod 𝓘(𝕜, V)) 𝓘(𝕜, V) ⊤ fun p : 𝕜 × V => p.1 • p.2 := + contMDiff_iff.2 ⟨continuous_smul, fun _ _ => contDiff_smul.contDiffOn⟩ + +@[deprecated (since := "2024-11-20")] alias smooth_smul := contMDiff_smul theorem ContMDiffWithinAt.smul {f : M → 𝕜} {g : M → V} (hf : ContMDiffWithinAt I 𝓘(𝕜) n f s x) (hg : ContMDiffWithinAt I 𝓘(𝕜, V) n g s x) : ContMDiffWithinAt I 𝓘(𝕜, V) n (fun p => f p • g p) s x := - (smooth_smul.of_le le_top).contMDiffAt.comp_contMDiffWithinAt x (hf.prod_mk hg) + (contMDiff_smul.of_le le_top).contMDiffAt.comp_contMDiffWithinAt x (hf.prod_mk hg) nonrec theorem ContMDiffAt.smul {f : M → 𝕜} {g : M → V} (hf : ContMDiffAt I 𝓘(𝕜) n f x) (hg : ContMDiffAt I 𝓘(𝕜, V) n g x) : ContMDiffAt I 𝓘(𝕜, V) n (fun p => f p • g p) x := @@ -281,18 +284,10 @@ theorem ContMDiff.smul {f : M → 𝕜} {g : M → V} (hf : ContMDiff I 𝓘( (hg : ContMDiff I 𝓘(𝕜, V) n g) : ContMDiff I 𝓘(𝕜, V) n fun p => f p • g p := fun x => (hf x).smul (hg x) -nonrec theorem SmoothWithinAt.smul {f : M → 𝕜} {g : M → V} (hf : SmoothWithinAt I 𝓘(𝕜) f s x) - (hg : SmoothWithinAt I 𝓘(𝕜, V) g s x) : SmoothWithinAt I 𝓘(𝕜, V) (fun p => f p • g p) s x := - hf.smul hg +@[deprecated (since := "2024-11-20")] alias SmoothWithinAt.smul := ContMDiffWithinAt.smul -nonrec theorem SmoothAt.smul {f : M → 𝕜} {g : M → V} (hf : SmoothAt I 𝓘(𝕜) f x) - (hg : SmoothAt I 𝓘(𝕜, V) g x) : SmoothAt I 𝓘(𝕜, V) (fun p => f p • g p) x := - hf.smul hg +@[deprecated (since := "2024-11-20")] alias SmoothAt.smul := ContMDiffAt.smul -nonrec theorem SmoothOn.smul {f : M → 𝕜} {g : M → V} (hf : SmoothOn I 𝓘(𝕜) f s) - (hg : SmoothOn I 𝓘(𝕜, V) g s) : SmoothOn I 𝓘(𝕜, V) (fun p => f p • g p) s := - hf.smul hg +@[deprecated (since := "2024-11-20")] alias SmoothOn.smul := ContMDiffOn.smul -nonrec theorem Smooth.smul {f : M → 𝕜} {g : M → V} (hf : Smooth I 𝓘(𝕜) f) - (hg : Smooth I 𝓘(𝕜, V) g) : Smooth I 𝓘(𝕜, V) fun p => f p • g p := - hf.smul hg +@[deprecated (since := "2024-11-20")] alias Smooth.smul := ContMDiff.smul diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Product.lean b/Mathlib/Geometry/Manifold/ContMDiff/Product.lean index 7e220374d4678..5562db2312500 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Product.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Product.lean @@ -81,38 +81,22 @@ theorem ContMDiff.prod_mk_space {f : M → E'} {g : M → F'} (hf : ContMDiff I (hg : ContMDiff I 𝓘(𝕜, F') n g) : ContMDiff I 𝓘(𝕜, E' × F') n fun x => (f x, g x) := fun x => (hf x).prod_mk_space (hg x) -nonrec theorem SmoothWithinAt.prod_mk {f : M → M'} {g : M → N'} (hf : SmoothWithinAt I I' f s x) - (hg : SmoothWithinAt I J' g s x) : SmoothWithinAt I (I'.prod J') (fun x => (f x, g x)) s x := - hf.prod_mk hg +@[deprecated (since := "2024-11-20")] alias SmoothWithinAt.prod_mk := ContMDiffWithinAt.prod_mk -nonrec theorem SmoothWithinAt.prod_mk_space {f : M → E'} {g : M → F'} - (hf : SmoothWithinAt I 𝓘(𝕜, E') f s x) (hg : SmoothWithinAt I 𝓘(𝕜, F') g s x) : - SmoothWithinAt I 𝓘(𝕜, E' × F') (fun x => (f x, g x)) s x := - hf.prod_mk_space hg +@[deprecated (since := "2024-11-20")] +alias SmoothWithinAt.prod_mk_space := ContMDiffWithinAt.prod_mk_space -nonrec theorem SmoothAt.prod_mk {f : M → M'} {g : M → N'} (hf : SmoothAt I I' f x) - (hg : SmoothAt I J' g x) : SmoothAt I (I'.prod J') (fun x => (f x, g x)) x := - hf.prod_mk hg +@[deprecated (since := "2024-11-20")] alias SmoothAt.prod_mk := ContMDiffAt.prod_mk -nonrec theorem SmoothAt.prod_mk_space {f : M → E'} {g : M → F'} (hf : SmoothAt I 𝓘(𝕜, E') f x) - (hg : SmoothAt I 𝓘(𝕜, F') g x) : SmoothAt I 𝓘(𝕜, E' × F') (fun x => (f x, g x)) x := - hf.prod_mk_space hg +@[deprecated (since := "2024-11-20")] alias SmoothAt.prod_mk_space := ContMDiffAt.prod_mk_space -nonrec theorem SmoothOn.prod_mk {f : M → M'} {g : M → N'} (hf : SmoothOn I I' f s) - (hg : SmoothOn I J' g s) : SmoothOn I (I'.prod J') (fun x => (f x, g x)) s := - hf.prod_mk hg +@[deprecated (since := "2024-11-20")] alias SmoothOn.prod_mk := ContMDiffOn.prod_mk -nonrec theorem SmoothOn.prod_mk_space {f : M → E'} {g : M → F'} (hf : SmoothOn I 𝓘(𝕜, E') f s) - (hg : SmoothOn I 𝓘(𝕜, F') g s) : SmoothOn I 𝓘(𝕜, E' × F') (fun x => (f x, g x)) s := - hf.prod_mk_space hg +@[deprecated (since := "2024-11-20")] alias SmoothOn.prod_mk_space := ContMDiffOn.prod_mk_space -nonrec theorem Smooth.prod_mk {f : M → M'} {g : M → N'} (hf : Smooth I I' f) (hg : Smooth I J' g) : - Smooth I (I'.prod J') fun x => (f x, g x) := - hf.prod_mk hg +@[deprecated (since := "2024-11-20")] alias Smooth.prod_mk := ContMDiff.prod_mk -nonrec theorem Smooth.prod_mk_space {f : M → E'} {g : M → F'} (hf : Smooth I 𝓘(𝕜, E') f) - (hg : Smooth I 𝓘(𝕜, F') g) : Smooth I 𝓘(𝕜, E' × F') fun x => (f x, g x) := - hf.prod_mk_space hg +@[deprecated (since := "2024-11-20")] alias Smooth.prod_mk_space := ContMDiff.prod_mk_space end ProdMk @@ -146,18 +130,13 @@ theorem contMDiffOn_fst {s : Set (M × N)} : ContMDiffOn (I.prod J) I n Prod.fst theorem contMDiff_fst : ContMDiff (I.prod J) I n (@Prod.fst M N) := fun _ => contMDiffAt_fst -theorem smoothWithinAt_fst {s : Set (M × N)} {p : M × N} : - SmoothWithinAt (I.prod J) I Prod.fst s p := - contMDiffWithinAt_fst +@[deprecated (since := "2024-11-20")] alias smoothWithinAt_fst := contMDiffWithinAt_fst -theorem smoothAt_fst {p : M × N} : SmoothAt (I.prod J) I Prod.fst p := - contMDiffAt_fst +@[deprecated (since := "2024-11-20")] alias smoothAt_fst := contMDiffAt_fst -theorem smoothOn_fst {s : Set (M × N)} : SmoothOn (I.prod J) I Prod.fst s := - contMDiffOn_fst +@[deprecated (since := "2024-11-20")] alias smoothOn_fst := contMDiffOn_fst -theorem smooth_fst : Smooth (I.prod J) I (@Prod.fst M N) := - contMDiff_fst +@[deprecated (since := "2024-11-20")] alias smooth_fst := contMDiff_fst theorem ContMDiffAt.fst {f : N → M × M'} {x : N} (hf : ContMDiffAt J (I.prod I') n f x) : ContMDiffAt J I n (fun x => (f x).1) x := @@ -167,12 +146,9 @@ theorem ContMDiff.fst {f : N → M × M'} (hf : ContMDiff J (I.prod I') n f) : ContMDiff J I n fun x => (f x).1 := contMDiff_fst.comp hf -theorem SmoothAt.fst {f : N → M × M'} {x : N} (hf : SmoothAt J (I.prod I') f x) : - SmoothAt J I (fun x => (f x).1) x := - smoothAt_fst.comp x hf +@[deprecated (since := "2024-11-20")] alias SmoothAt.fst := ContMDiffAt.fst -theorem Smooth.fst {f : N → M × M'} (hf : Smooth J (I.prod I') f) : Smooth J I fun x => (f x).1 := - smooth_fst.comp hf +@[deprecated (since := "2024-11-20")] alias Smooth.fst := ContMDiff.fst theorem contMDiffWithinAt_snd {s : Set (M × N)} {p : M × N} : ContMDiffWithinAt (I.prod J) J n Prod.snd s p := by @@ -202,18 +178,13 @@ theorem contMDiffOn_snd {s : Set (M × N)} : ContMDiffOn (I.prod J) J n Prod.snd theorem contMDiff_snd : ContMDiff (I.prod J) J n (@Prod.snd M N) := fun _ => contMDiffAt_snd -theorem smoothWithinAt_snd {s : Set (M × N)} {p : M × N} : - SmoothWithinAt (I.prod J) J Prod.snd s p := - contMDiffWithinAt_snd +@[deprecated (since := "2024-11-20")] alias smoothWithinAt_snd := contMDiffWithinAt_snd -theorem smoothAt_snd {p : M × N} : SmoothAt (I.prod J) J Prod.snd p := - contMDiffAt_snd +@[deprecated (since := "2024-11-20")] alias smoothAt_snd := contMDiffAt_snd -theorem smoothOn_snd {s : Set (M × N)} : SmoothOn (I.prod J) J Prod.snd s := - contMDiffOn_snd +@[deprecated (since := "2024-11-20")] alias smoothOn_snd := contMDiffOn_snd -theorem smooth_snd : Smooth (I.prod J) J (@Prod.snd M N) := - contMDiff_snd +@[deprecated (since := "2024-11-20")] alias smooth_snd := contMDiff_snd theorem ContMDiffAt.snd {f : N → M × M'} {x : N} (hf : ContMDiffAt J (I.prod I') n f x) : ContMDiffAt J I' n (fun x => (f x).2) x := @@ -223,12 +194,9 @@ theorem ContMDiff.snd {f : N → M × M'} (hf : ContMDiff J (I.prod I') n f) : ContMDiff J I' n fun x => (f x).2 := contMDiff_snd.comp hf -theorem SmoothAt.snd {f : N → M × M'} {x : N} (hf : SmoothAt J (I.prod I') f x) : - SmoothAt J I' (fun x => (f x).2) x := - smoothAt_snd.comp x hf +@[deprecated (since := "2024-11-20")] alias SmoothAt.snd := ContMDiffAt.snd -theorem Smooth.snd {f : N → M × M'} (hf : Smooth J (I.prod I') f) : Smooth J I' fun x => (f x).2 := - smooth_snd.comp hf +@[deprecated (since := "2024-11-20")] alias Smooth.snd := ContMDiff.snd end Projections @@ -279,17 +247,16 @@ theorem contMDiff_prod_module_iff (f : M → F₁ × F₂) : rw [modelWithCornersSelf_prod, ← chartedSpaceSelf_prod] exact contMDiff_prod_iff f -theorem smoothAt_prod_iff (f : M → M' × N') {x : M} : - SmoothAt I (I'.prod J') f x ↔ SmoothAt I I' (Prod.fst ∘ f) x ∧ SmoothAt I J' (Prod.snd ∘ f) x := - contMDiffAt_prod_iff f +theorem contMDiff_prod_assoc : + ContMDiff ((I.prod I').prod J) (I.prod (I'.prod J)) n + fun x : (M × M') × N => (x.1.1, x.1.2, x.2) := + contMDiff_fst.fst.prod_mk <| contMDiff_fst.snd.prod_mk contMDiff_snd + +@[deprecated (since := "2024-11-20")] alias smoothAt_prod_iff := contMDiffAt_prod_iff -theorem smooth_prod_iff (f : M → M' × N') : - Smooth I (I'.prod J') f ↔ Smooth I I' (Prod.fst ∘ f) ∧ Smooth I J' (Prod.snd ∘ f) := - contMDiff_prod_iff f +@[deprecated (since := "2024-11-20")] alias smooth_prod_iff := contMDiff_prod_iff -theorem smooth_prod_assoc : - Smooth ((I.prod I').prod J) (I.prod (I'.prod J)) fun x : (M × M') × N => (x.1.1, x.1.2, x.2) := - smooth_fst.fst.prod_mk <| smooth_fst.snd.prod_mk smooth_snd +@[deprecated (since := "2024-11-20")] alias smooth_prod_assoc := contMDiff_prod_assoc section prodMap @@ -329,22 +296,13 @@ theorem ContMDiff.prod_map (hf : ContMDiff I I' n f) (hg : ContMDiff J J' n g) : intro p exact (hf p.1).prod_map' (hg p.2) -nonrec theorem SmoothWithinAt.prod_map (hf : SmoothWithinAt I I' f s x) - (hg : SmoothWithinAt J J' g r y) : - SmoothWithinAt (I.prod J) (I'.prod J') (Prod.map f g) (s ×ˢ r) (x, y) := - hf.prod_map hg +@[deprecated (since := "2024-11-20")] alias SmoothWithinAt.prod_map := ContMDiffWithinAt.prod_map -nonrec theorem SmoothAt.prod_map (hf : SmoothAt I I' f x) (hg : SmoothAt J J' g y) : - SmoothAt (I.prod J) (I'.prod J') (Prod.map f g) (x, y) := - hf.prod_map hg +@[deprecated (since := "2024-11-20")] alias SmoothAt.prod_map := ContMDiffAt.prod_map -nonrec theorem SmoothOn.prod_map (hf : SmoothOn I I' f s) (hg : SmoothOn J J' g r) : - SmoothOn (I.prod J) (I'.prod J') (Prod.map f g) (s ×ˢ r) := - hf.prod_map hg +@[deprecated (since := "2024-11-20")] alias SmoothOn.prod_map := ContMDiffOn.prod_map -nonrec theorem Smooth.prod_map (hf : Smooth I I' f) (hg : Smooth J J' g) : - Smooth (I.prod J) (I'.prod J') (Prod.map f g) := - hf.prod_map hg +@[deprecated (since := "2024-11-20")] alias Smooth.prod_map := ContMDiff.prod_map end prodMap @@ -380,20 +338,12 @@ theorem contMDiff_pi_space : ContMDiff I 𝓘(𝕜, ∀ i, Fi i) n φ ↔ ∀ i, ContMDiff I 𝓘(𝕜, Fi i) n fun x => φ x i := ⟨fun h i x => contMDiffAt_pi_space.1 (h x) i, fun h x => contMDiffAt_pi_space.2 fun i => h i x⟩ -theorem smoothWithinAt_pi_space : - SmoothWithinAt I 𝓘(𝕜, ∀ i, Fi i) φ s x ↔ - ∀ i, SmoothWithinAt I 𝓘(𝕜, Fi i) (fun x => φ x i) s x := - contMDiffWithinAt_pi_space +@[deprecated (since := "2024-11-20")] alias smoothWithinAt_pi_space := contMDiffWithinAt_pi_space -theorem smoothOn_pi_space : - SmoothOn I 𝓘(𝕜, ∀ i, Fi i) φ s ↔ ∀ i, SmoothOn I 𝓘(𝕜, Fi i) (fun x => φ x i) s := - contMDiffOn_pi_space +@[deprecated (since := "2024-11-20")] alias smoothAt_pi_space := contMDiffAt_pi_space -theorem smoothAt_pi_space : - SmoothAt I 𝓘(𝕜, ∀ i, Fi i) φ x ↔ ∀ i, SmoothAt I 𝓘(𝕜, Fi i) (fun x => φ x i) x := - contMDiffAt_pi_space +@[deprecated (since := "2024-11-20")] alias smoothOn_pi_space := contMDiffOn_pi_space -theorem smooth_pi_space : Smooth I 𝓘(𝕜, ∀ i, Fi i) φ ↔ ∀ i, Smooth I 𝓘(𝕜, Fi i) fun x => φ x i := - contMDiff_pi_space +@[deprecated (since := "2024-11-20")] alias smooth_pi_space := contMDiff_pi_space end PiSpace diff --git a/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean b/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean index 9ab1a8b787d06..ad3c237062762 100644 --- a/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean +++ b/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean @@ -343,10 +343,7 @@ theorem ContMDiff.continuous_tangentMap (hf : ContMDiff I I' n f) (hmn : 1 ≤ n convert hf.continuousOn_tangentMapWithin hmn uniqueMDiffOn_univ rw [tangentMapWithin_univ] -/-- If a function is smooth, then its bundled derivative is smooth. -/ -theorem Smooth.tangentMap (hf : Smooth I I' f) : - Smooth I.tangent I'.tangent (tangentMap I I' f) := - ContMDiff.contMDiff_tangentMap hf le_rfl +@[deprecated (since := "2024-11-21")] alias Smooth.tangentMap := ContMDiff.contMDiff_tangentMap end tangentMap @@ -376,9 +373,9 @@ theorem tangentMap_tangentBundle_pure [Is : SmoothManifoldWithCorners I M] (p : · apply (PartialHomeomorph.open_target _).preimage I.continuous_invFun · simp only [mfld_simps] have A : MDifferentiableAt I I.tangent (fun x => @TotalSpace.mk M E (TangentSpace I) x 0) x := - haveI : Smooth I (I.prod 𝓘(𝕜, E)) (zeroSection E (TangentSpace I : M → Type _)) := - Bundle.smooth_zeroSection 𝕜 (TangentSpace I : M → Type _) - this.mdifferentiableAt + haveI : ContMDiff I (I.prod 𝓘(𝕜, E)) ⊤ (zeroSection E (TangentSpace I : M → Type _)) := + Bundle.contMDiff_zeroSection 𝕜 (TangentSpace I : M → Type _) + this.mdifferentiableAt le_top have B : fderivWithin 𝕜 (fun x' : E ↦ (x', (0 : E))) (Set.range I) (I ((chartAt H x) x)) v = (v, 0) := by rw [fderivWithin_eq_fderiv, DifferentiableAt.fderiv_prod] diff --git a/Mathlib/Geometry/Manifold/ContMDiffMap.lean b/Mathlib/Geometry/Manifold/ContMDiffMap.lean index 8566926007fa4..f8bd9f7980ad8 100644 --- a/Mathlib/Geometry/Manifold/ContMDiffMap.lean +++ b/Mathlib/Geometry/Manifold/ContMDiffMap.lean @@ -29,10 +29,7 @@ variable (I I') in def ContMDiffMap := { f : M → M' // ContMDiff I I' n f } -variable (I I') in -/-- Bundled smooth maps. -/ -abbrev SmoothMap := - ContMDiffMap I I' M M' ⊤ +@[deprecated (since := "024-11-21")] alias SmoothMap := ContMDiffMap @[inherit_doc] scoped[Manifold] notation "C^" n "⟮" I ", " M "; " I' ", " M' "⟯" => ContMDiffMap I I' M M' n @@ -54,8 +51,7 @@ instance instFunLike : FunLike C^n⟮I, M; I', M'⟯ M M' where protected theorem contMDiff (f : C^n⟮I, M; I', M'⟯) : ContMDiff I I' n f := f.prop -protected theorem smooth (f : C^∞⟮I, M; I', M'⟯) : Smooth I I' f := - f.prop +@[deprecated (since := "2024-11-20")] alias smooth := ContMDiffMap.contMDiff -- Porting note: use generic instance instead -- instance : Coe C^n⟮I, M; I', M'⟯ C(M, M') := diff --git a/Mathlib/Geometry/Manifold/Diffeomorph.lean b/Mathlib/Geometry/Manifold/Diffeomorph.lean index 8a2f3e373a09a..5ee61a8605fe2 100644 --- a/Mathlib/Geometry/Manifold/Diffeomorph.lean +++ b/Mathlib/Geometry/Manifold/Diffeomorph.lean @@ -125,7 +125,7 @@ protected theorem contMDiffWithinAt (h : M ≃ₘ^n⟮I, I'⟯ M') {s x} : ContM protected theorem contDiff (h : E ≃ₘ^n⟮𝓘(𝕜, E), 𝓘(𝕜, E')⟯ E') : ContDiff 𝕜 n h := h.contMDiff.contDiff -protected theorem smooth (h : M ≃ₘ⟮I, I'⟯ M') : Smooth I I' h := h.contMDiff +@[deprecated (since := "2024-11-21")] alias smooth := Diffeomorph.contDiff protected theorem mdifferentiable (h : M ≃ₘ^n⟮I, I'⟯ M') (hn : 1 ≤ n) : MDifferentiable I I' h := h.contMDiff.mdifferentiable hn @@ -539,9 +539,8 @@ theorem contMDiff_transDiffeomorph_right {f : M' → M} : ContMDiff I' (I.transDiffeomorph e) n f ↔ ContMDiff I' I n f := (toTransDiffeomorph I M e).contMDiff_diffeomorph_comp_iff le_top -theorem smooth_transDiffeomorph_right {f : M' → M} : - Smooth I' (I.transDiffeomorph e) f ↔ Smooth I' I f := - contMDiff_transDiffeomorph_right e +@[deprecated (since := "2024-11-21")] +alias smooth_transDiffeomorph_right := contMDiff_transDiffeomorph_right @[simp] theorem contMDiffWithinAt_transDiffeomorph_left {f : M → M'} {x s} : @@ -563,8 +562,7 @@ theorem contMDiff_transDiffeomorph_left {f : M → M'} : ContMDiff (I.transDiffeomorph e) I' n f ↔ ContMDiff I I' n f := ((toTransDiffeomorph I M e).contMDiff_comp_diffeomorph_iff le_top).symm -theorem smooth_transDiffeomorph_left {f : M → M'} : - Smooth (I.transDiffeomorph e) I' f ↔ Smooth I I' f := - e.contMDiff_transDiffeomorph_left +@[deprecated (since := "2024-11-21")] +alias smooth_transDiffeomorph_left := contMDiff_transDiffeomorph_left end Diffeomorph diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean index 552b4986092f8..377ab9d17b795 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean @@ -477,27 +477,30 @@ theorem ContMDiff.mdifferentiableAt (hf : ContMDiff I I' n f) (hn : 1 ≤ n) : MDifferentiableAt I I' f x := hf.contMDiffAt.mdifferentiableAt hn +theorem ContMDiff.mdifferentiableWithinAt (hf : ContMDiff I I' n f) (hn : 1 ≤ n) : + MDifferentiableWithinAt I I' f s x := + (hf.contMDiffAt.mdifferentiableAt hn).mdifferentiableWithinAt + theorem ContMDiffOn.mdifferentiableOn (hf : ContMDiffOn I I' n f s) (hn : 1 ≤ n) : MDifferentiableOn I I' f s := fun x hx => (hf x hx).mdifferentiableWithinAt hn -nonrec theorem SmoothWithinAt.mdifferentiableWithinAt (hf : SmoothWithinAt I I' f s x) : - MDifferentiableWithinAt I I' f s x := - hf.mdifferentiableWithinAt le_top +@[deprecated (since := "2024-11-20")] +alias SmoothWithinAt.mdifferentiableWithinAt := ContMDiffWithinAt.mdifferentiableWithinAt theorem ContMDiff.mdifferentiable (hf : ContMDiff I I' n f) (hn : 1 ≤ n) : MDifferentiable I I' f := fun x => (hf x).mdifferentiableAt hn -nonrec theorem SmoothAt.mdifferentiableAt (hf : SmoothAt I I' f x) : MDifferentiableAt I I' f x := - hf.mdifferentiableAt le_top +@[deprecated (since := "2024-11-20")] +alias SmoothAt.mdifferentiableAt := ContMDiffAt.mdifferentiableAt -nonrec theorem SmoothOn.mdifferentiableOn (hf : SmoothOn I I' f s) : MDifferentiableOn I I' f s := - hf.mdifferentiableOn le_top +@[deprecated (since := "2024-11-20")] +alias SmoothOn.mdifferentiableOn := ContMDiffOn.mdifferentiableOn -theorem Smooth.mdifferentiable (hf : Smooth I I' f) : MDifferentiable I I' f := - ContMDiff.mdifferentiable hf le_top +@[deprecated (since := "2024-11-20")] +alias Smooth.mdifferentiable := ContMDiff.mdifferentiable -theorem Smooth.mdifferentiableAt (hf : Smooth I I' f) : MDifferentiableAt I I' f x := - hf.mdifferentiable x +@[deprecated (since := "2024-11-20")] +alias Smooth.mdifferentiableAt := ContMDiff.mdifferentiableAt theorem MDifferentiableOn.continuousOn (h : MDifferentiableOn I I' f s) : ContinuousOn f s := fun x hx => (h x hx).continuousWithinAt @@ -505,8 +508,8 @@ theorem MDifferentiableOn.continuousOn (h : MDifferentiableOn I I' f s) : Contin theorem MDifferentiable.continuous (h : MDifferentiable I I' f) : Continuous f := continuous_iff_continuousAt.2 fun x => (h x).continuousAt -theorem Smooth.mdifferentiableWithinAt (hf : Smooth I I' f) : MDifferentiableWithinAt I I' f s x := - hf.mdifferentiableAt.mdifferentiableWithinAt +@[deprecated (since := "2024-11-20")] +alias Smooth.mdifferentiableWithinAt := ContMDiff.mdifferentiableWithinAt /-! ### Deriving continuity from differentiability on manifolds -/ diff --git a/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean b/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean index 4facc7f820910..e090b101ab9e7 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean @@ -120,7 +120,7 @@ variable {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] {Z : M → Type theorem Trivialization.mdifferentiable (e : Trivialization F (π F Z)) [MemTrivializationAtlas e] : e.toPartialHomeomorph.MDifferentiable (I.prod 𝓘(𝕜, F)) (I.prod 𝓘(𝕜, F)) := - ⟨e.smoothOn.mdifferentiableOn, e.smoothOn_symm.mdifferentiableOn⟩ + ⟨e.contMDiffOn.mdifferentiableOn le_top, e.contMDiffOn_symm.mdifferentiableOn le_top⟩ theorem UniqueMDiffWithinAt.smooth_bundle_preimage {p : TotalSpace F Z} (hs : UniqueMDiffWithinAt I s p.proj) : diff --git a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean index 35ffa8cfc8743..3d001e4e01c5d 100644 --- a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean +++ b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean @@ -160,8 +160,10 @@ theorem sum_le_one (x : M) : ∑ᶠ i, f i x ≤ 1 := def toPartitionOfUnity : PartitionOfUnity ι M s := { f with toFun := fun i => f i } -theorem smooth_sum : Smooth I 𝓘(ℝ) fun x => ∑ᶠ i, f i x := - smooth_finsum (fun i => (f i).smooth) f.locallyFinite +theorem contMDiff_sum : ContMDiff I 𝓘(ℝ) ⊤ fun x => ∑ᶠ i, f i x := + contMDiff_finsum (fun i => (f i).contMDiff) f.locallyFinite + +@[deprecated (since := "2024-11-21")] alias smooth_sum := contMDiff_sum theorem le_one (i : ι) (x : M) : f i x ≤ 1 := f.toPartitionOfUnity.le_one i x @@ -178,9 +180,7 @@ theorem contMDiff_smul {g : M → F} {i} (hg : ∀ x ∈ tsupport (f i), ContMDi contMDiff_of_tsupport fun x hx => ((f i).contMDiff.contMDiffAt.of_le le_top).smul <| hg x <| tsupport_smul_subset_left _ _ hx -theorem smooth_smul {g : M → F} {i} (hg : ∀ x ∈ tsupport (f i), SmoothAt I 𝓘(ℝ, F) g x) : - Smooth I 𝓘(ℝ, F) fun x => f i x • g x := - f.contMDiff_smul hg +@[deprecated (since := "2024-11-21")] alias smooth_smul := contMDiff_smul /-- If `f` is a smooth partition of unity on a set `s : Set M` and `g : ι → M → F` is a family of functions such that `g i` is $C^n$ smooth at every point of the topological support of `f i`, then @@ -191,20 +191,14 @@ theorem contMDiff_finsum_smul {g : ι → M → F} (contMDiff_finsum fun i => f.contMDiff_smul (hg i)) <| f.locallyFinite.subset fun _ => support_smul_subset_left _ _ -/-- If `f` is a smooth partition of unity on a set `s : Set M` and `g : ι → M → F` is a family of -functions such that `g i` is smooth at every point of the topological support of `f i`, then the sum -`fun x ↦ ∑ᶠ i, f i x • g i x` is smooth on the whole manifold. -/ -theorem smooth_finsum_smul {g : ι → M → F} - (hg : ∀ (i), ∀ x ∈ tsupport (f i), SmoothAt I 𝓘(ℝ, F) (g i) x) : - Smooth I 𝓘(ℝ, F) fun x => ∑ᶠ i, f i x • g i x := - f.contMDiff_finsum_smul hg +@[deprecated (since := "2024-11-21")] alias smooth_finsum_smul := contMDiff_finsum_smul theorem contMDiffAt_finsum {x₀ : M} {g : ι → M → F} (hφ : ∀ i, x₀ ∈ tsupport (f i) → ContMDiffAt I 𝓘(ℝ, F) n (g i) x₀) : ContMDiffAt I 𝓘(ℝ, F) n (fun x ↦ ∑ᶠ i, f i x • g i x) x₀ := by refine _root_.contMDiffAt_finsum (f.locallyFinite.smul_left _) fun i ↦ ?_ by_cases hx : x₀ ∈ tsupport (f i) - · exact ContMDiffAt.smul ((f i).smooth.of_le le_top).contMDiffAt (hφ i hx) + · exact ContMDiffAt.smul ((f i).contMDiff.of_le le_top).contMDiffAt (hφ i hx) · exact contMDiffAt_of_not_mem (compl_subset_compl.mpr (tsupport_smul_subset_left (f i) (g i)) hx) n @@ -294,13 +288,8 @@ theorem IsSubordinate.contMDiff_finsum_smul {g : ι → M → F} (hf : f.IsSubor ContMDiff I 𝓘(ℝ, F) n fun x => ∑ᶠ i, f i x • g i x := f.contMDiff_finsum_smul fun i _ hx => (hg i).contMDiffAt <| (ho i).mem_nhds (hf i hx) -/-- If `f` is a smooth partition of unity on a set `s : Set M` subordinate to a family of open sets -`U : ι → Set M` and `g : ι → M → F` is a family of functions such that `g i` is smooth on `U i`, -then the sum `fun x ↦ ∑ᶠ i, f i x • g i x` is smooth on the whole manifold. -/ -theorem IsSubordinate.smooth_finsum_smul {g : ι → M → F} (hf : f.IsSubordinate U) - (ho : ∀ i, IsOpen (U i)) (hg : ∀ i, SmoothOn I 𝓘(ℝ, F) (g i) (U i)) : - Smooth I 𝓘(ℝ, F) fun x => ∑ᶠ i, f i x • g i x := - hf.contMDiff_finsum_smul ho hg +@[deprecated (since := "2024-11-21")] +alias IsSubordinate.smooth_finsum_smul := IsSubordinate.contMDiff_finsum_smul end IsSubordinate @@ -309,14 +298,17 @@ end SmoothPartitionOfUnity namespace BumpCovering -- Repeat variables to drop `[FiniteDimensional ℝ E]` and `[SmoothManifoldWithCorners I M]` -theorem smooth_toPartitionOfUnity {E : Type uE} [NormedAddCommGroup E] [NormedSpace ℝ E] +theorem contMDiff_toPartitionOfUnity {E : Type uE} [NormedAddCommGroup E] [NormedSpace ℝ E] {H : Type uH} [TopologicalSpace H] {I : ModelWithCorners ℝ E H} {M : Type uM} [TopologicalSpace M] [ChartedSpace H M] {s : Set M} (f : BumpCovering ι M s) - (hf : ∀ i, Smooth I 𝓘(ℝ) (f i)) (i : ι) : Smooth I 𝓘(ℝ) (f.toPartitionOfUnity i) := - (hf i).mul <| (smooth_finprod_cond fun j _ => smooth_const.sub (hf j)) <| by + (hf : ∀ i, ContMDiff I 𝓘(ℝ) ⊤ (f i)) (i : ι) : ContMDiff I 𝓘(ℝ) ⊤ (f.toPartitionOfUnity i) := + (hf i).mul <| (contMDiff_finprod_cond fun j _ => contMDiff_const.sub (hf j)) <| by simp only [Pi.sub_def, mulSupport_one_sub] exact f.locallyFinite +@[deprecated (since := "2024-11-21")] +alias smooth_toPartitionOfUnity := contMDiff_toPartitionOfUnity + variable {s : Set M} /-- A `BumpCovering` such that all functions in this covering are smooth generates a smooth @@ -326,24 +318,24 @@ In our formalization, not every `f : BumpCovering ι M s` with smooth functions `SmoothBumpCovering`; instead, a `SmoothBumpCovering` is a covering by supports of `SmoothBumpFunction`s. So, we define `BumpCovering.toSmoothPartitionOfUnity`, then reuse it in `SmoothBumpCovering.toSmoothPartitionOfUnity`. -/ -def toSmoothPartitionOfUnity (f : BumpCovering ι M s) (hf : ∀ i, Smooth I 𝓘(ℝ) (f i)) : +def toSmoothPartitionOfUnity (f : BumpCovering ι M s) (hf : ∀ i, ContMDiff I 𝓘(ℝ) ⊤ (f i)) : SmoothPartitionOfUnity ι I M s := { f.toPartitionOfUnity with - toFun := fun i => ⟨f.toPartitionOfUnity i, f.smooth_toPartitionOfUnity hf i⟩ } + toFun := fun i => ⟨f.toPartitionOfUnity i, f.contMDiff_toPartitionOfUnity hf i⟩ } @[simp] theorem toSmoothPartitionOfUnity_toPartitionOfUnity (f : BumpCovering ι M s) - (hf : ∀ i, Smooth I 𝓘(ℝ) (f i)) : + (hf : ∀ i, ContMDiff I 𝓘(ℝ) ⊤ (f i)) : (f.toSmoothPartitionOfUnity hf).toPartitionOfUnity = f.toPartitionOfUnity := rfl @[simp] -theorem coe_toSmoothPartitionOfUnity (f : BumpCovering ι M s) (hf : ∀ i, Smooth I 𝓘(ℝ) (f i)) +theorem coe_toSmoothPartitionOfUnity (f : BumpCovering ι M s) (hf : ∀ i, ContMDiff I 𝓘(ℝ) ⊤ (f i)) (i : ι) : ⇑(f.toSmoothPartitionOfUnity hf i) = f.toPartitionOfUnity i := rfl theorem IsSubordinate.toSmoothPartitionOfUnity {f : BumpCovering ι M s} {U : ι → Set M} - (h : f.IsSubordinate U) (hf : ∀ i, Smooth I 𝓘(ℝ) (f i)) : + (h : f.IsSubordinate U) (hf : ∀ i, ContMDiff I 𝓘(ℝ) ⊤ (f i)) : (f.toSmoothPartitionOfUnity hf).IsSubordinate U := h.toPartitionOfUnity @@ -457,7 +449,7 @@ alias ⟨_, IsSubordinate.toBumpCovering⟩ := isSubordinate_toBumpCovering /-- Every `SmoothBumpCovering` defines a smooth partition of unity. -/ def toSmoothPartitionOfUnity : SmoothPartitionOfUnity ι I M s := - fs.toBumpCovering.toSmoothPartitionOfUnity fun i => (fs i).smooth + fs.toBumpCovering.toSmoothPartitionOfUnity fun i => (fs i).contMDiff theorem toSmoothPartitionOfUnity_apply (i : ι) (x : M) : fs.toSmoothPartitionOfUnity i x = fs i x * ∏ᶠ (j) (_ : WellOrderingRel j i), (1 - fs j x) := @@ -511,7 +503,7 @@ theorem exists_smooth_zero_one_of_isClosed [T2Space M] [SigmaCompactSpace M] {s rcases SmoothBumpCovering.exists_isSubordinate I ht this with ⟨ι, f, hf⟩ set g := f.toSmoothPartitionOfUnity refine - ⟨⟨_, g.smooth_sum⟩, fun x hx => ?_, fun x => g.sum_eq_one, fun x => + ⟨⟨_, g.contMDiff_sum⟩, fun x hx => ?_, fun x => g.sum_eq_one, fun x => ⟨g.sum_nonneg x, g.sum_le_one x⟩⟩ suffices ∀ i, g i x = 0 by simp only [this, ContMDiffMap.coeFn_mk, finsum_zero, Pi.zero_apply] refine fun i => f.toSmoothPartitionOfUnity_zero_of_zero ?_ @@ -554,8 +546,9 @@ def single (i : ι) (s : Set M) : SmoothPartitionOfUnity ι I M s := (BumpCovering.single i s).toSmoothPartitionOfUnity fun j => by classical rcases eq_or_ne j i with (rfl | h) - · simp only [smooth_one, ContinuousMap.coe_one, BumpCovering.coe_single, Pi.single_eq_same] - · simp only [smooth_zero, BumpCovering.coe_single, Pi.single_eq_of_ne h, ContinuousMap.coe_zero] + · simp only [contMDiff_one, ContinuousMap.coe_one, BumpCovering.coe_single, Pi.single_eq_same] + · simp only [contMDiff_zero, BumpCovering.coe_single, Pi.single_eq_of_ne h, + ContinuousMap.coe_zero] instance [Inhabited ι] (s : Set M) : Inhabited (SmoothPartitionOfUnity ι I M s) := ⟨single I default s⟩ @@ -570,12 +563,12 @@ theorem exists_isSubordinate {s : Set M} (hs : IsClosed s) (U : ι → Set M) (h haveI : LocallyCompactSpace M := ChartedSpace.locallyCompactSpace H M -- porting note(https://github.com/leanprover/std4/issues/116): -- split `rcases` into `have` + `rcases` - have := BumpCovering.exists_isSubordinate_of_prop (Smooth I 𝓘(ℝ)) ?_ hs U ho hU + have := BumpCovering.exists_isSubordinate_of_prop (ContMDiff I 𝓘(ℝ) ⊤) ?_ hs U ho hU · rcases this with ⟨f, hf, hfU⟩ exact ⟨f.toSmoothPartitionOfUnity hf, hfU.toSmoothPartitionOfUnity hf⟩ · intro s t hs ht hd rcases exists_smooth_zero_one_of_isClosed I hs ht hd with ⟨f, hf⟩ - exact ⟨f, f.smooth, hf⟩ + exact ⟨f, f.contMDiff, hf⟩ theorem exists_isSubordinate_chartAt_source_of_isClosed {s : Set M} (hs : IsClosed s) : ∃ f : SmoothPartitionOfUnity s I M s, @@ -618,7 +611,7 @@ Then there exists a smooth function `g : C^∞⟮I, M; 𝓘(ℝ, F), F⟯` such See also `exists_contMDiffOn_forall_mem_convex_of_local` and `exists_smooth_forall_mem_convex_of_local_const`. -/ theorem exists_smooth_forall_mem_convex_of_local (ht : ∀ x, Convex ℝ (t x)) - (Hloc : ∀ x : M, ∃ U ∈ 𝓝 x, ∃ g : M → F, SmoothOn I 𝓘(ℝ, F) g U ∧ ∀ y ∈ U, g y ∈ t y) : + (Hloc : ∀ x : M, ∃ U ∈ 𝓝 x, ∃ g : M → F, ContMDiffOn I 𝓘(ℝ, F) ⊤ g U ∧ ∀ y ∈ U, g y ∈ t y) : ∃ g : C^∞⟮I, M; 𝓘(ℝ, F), F⟯, ∀ x, g x ∈ t x := exists_contMDiffOn_forall_mem_convex_of_local I ht Hloc @@ -631,7 +624,7 @@ theorem exists_smooth_forall_mem_convex_of_local_const (ht : ∀ x, Convex ℝ ( (Hloc : ∀ x : M, ∃ c : F, ∀ᶠ y in 𝓝 x, c ∈ t y) : ∃ g : C^∞⟮I, M; 𝓘(ℝ, F), F⟯, ∀ x, g x ∈ t x := exists_smooth_forall_mem_convex_of_local I ht fun x => let ⟨c, hc⟩ := Hloc x - ⟨_, hc, fun _ => c, smoothOn_const, fun _ => id⟩ + ⟨_, hc, fun _ => c, contMDiffOn_const, fun _ => id⟩ /-- Let `M` be a smooth σ-compact manifold with extended distance. Let `K : ι → Set M` be a locally finite family of closed sets, let `U : ι → Set M` be a family of open sets such that `K i ⊆ U i` for @@ -664,7 +657,7 @@ theorem Metric.exists_smooth_forall_closedBall_subset {M} [MetricSpace M] [Chart exact hδ i x hx lemma IsOpen.exists_msmooth_support_eq_aux {s : Set H} (hs : IsOpen s) : - ∃ f : H → ℝ, f.support = s ∧ Smooth I 𝓘(ℝ) f ∧ Set.range f ⊆ Set.Icc 0 1 := by + ∃ f : H → ℝ, f.support = s ∧ ContMDiff I 𝓘(ℝ) ⊤ f ∧ Set.range f ⊆ Set.Icc 0 1 := by have h's : IsOpen (I.symm ⁻¹' s) := I.continuous_symm.isOpen_preimage _ hs rcases h's.exists_smooth_support_eq with ⟨f, f_supp, f_diff, f_range⟩ refine ⟨f ∘ I, ?_, ?_, ?_⟩ @@ -676,11 +669,11 @@ lemma IsOpen.exists_msmooth_support_eq_aux {s : Set H} (hs : IsOpen s) : /-- Given an open set in a finite-dimensional real manifold, there exists a nonnegative smooth function with support equal to `s`. -/ theorem IsOpen.exists_msmooth_support_eq {s : Set M} (hs : IsOpen s) : - ∃ f : M → ℝ, f.support = s ∧ Smooth I 𝓘(ℝ) f ∧ ∀ x, 0 ≤ f x := by + ∃ f : M → ℝ, f.support = s ∧ ContMDiff I 𝓘(ℝ) ⊤ f ∧ ∀ x, 0 ≤ f x := by rcases SmoothPartitionOfUnity.exists_isSubordinate_chartAt_source I M with ⟨f, hf⟩ have A : ∀ (c : M), ∃ g : H → ℝ, g.support = (chartAt H c).target ∩ (chartAt H c).symm ⁻¹' s ∧ - Smooth I 𝓘(ℝ) g ∧ Set.range g ⊆ Set.Icc 0 1 := by + ContMDiff I 𝓘(ℝ) ⊤ g ∧ Set.range g ⊆ Set.Icc 0 1 := by intro i apply IsOpen.exists_msmooth_support_eq_aux exact PartialHomeomorph.isOpen_inter_preimage_symm _ hs @@ -714,7 +707,7 @@ theorem IsOpen.exists_msmooth_support_eq {s : Set M} (hs : IsOpen s) : · have : x ∉ support (f c) := by contrapose! Hx; exact subset_tsupport _ Hx rw [nmem_support] at this simp [this] - · apply SmoothPartitionOfUnity.smooth_finsum_smul + · apply SmoothPartitionOfUnity.contMDiff_finsum_smul intro c x hx apply (g_diff c (chartAt H c x)).comp exact contMDiffAt_of_mem_maximalAtlas (SmoothManifoldWithCorners.chart_mem_maximalAtlas _) @@ -727,7 +720,7 @@ exists a smooth function with support equal to `s`, taking values in `[0,1]`, an exactly on `t`. -/ theorem exists_msmooth_support_eq_eq_one_iff {s t : Set M} (hs : IsOpen s) (ht : IsClosed t) (h : t ⊆ s) : - ∃ f : M → ℝ, Smooth I 𝓘(ℝ) f ∧ range f ⊆ Icc 0 1 ∧ support f = s + ∃ f : M → ℝ, ContMDiff I 𝓘(ℝ) ⊤ f ∧ range f ⊆ Icc 0 1 ∧ support f = s ∧ (∀ x, x ∈ t ↔ f x = 1) := by /- Take `f` with support equal to `s`, and `g` with support equal to `tᶜ`. Then `f / (f + g)` satisfies the conclusion of the theorem. -/ @@ -765,7 +758,7 @@ there exists an infinitely smooth function that is equal to `0` exactly on `s` a exactly on `t`. See also `exists_smooth_zero_one_of_isClosed` for a slightly weaker version. -/ theorem exists_msmooth_zero_iff_one_iff_of_isClosed {s t : Set M} (hs : IsClosed s) (ht : IsClosed t) (hd : Disjoint s t) : - ∃ f : M → ℝ, Smooth I 𝓘(ℝ) f ∧ range f ⊆ Icc 0 1 ∧ (∀ x, x ∈ s ↔ f x = 0) + ∃ f : M → ℝ, ContMDiff I 𝓘(ℝ) ⊤ f ∧ range f ⊆ Icc 0 1 ∧ (∀ x, x ∈ s ↔ f x = 0) ∧ (∀ x, x ∈ t ↔ f x = 1) := by rcases exists_msmooth_support_eq_eq_one_iff I hs.isOpen_compl ht hd.subset_compl_left with ⟨f, f_diff, f_range, fs, ft⟩ diff --git a/Mathlib/Geometry/Manifold/Sheaf/LocallyRingedSpace.lean b/Mathlib/Geometry/Manifold/Sheaf/LocallyRingedSpace.lean index 889506e8609ea..48aec60a959fa 100644 --- a/Mathlib/Geometry/Manifold/Sheaf/LocallyRingedSpace.lean +++ b/Mathlib/Geometry/Manifold/Sheaf/LocallyRingedSpace.lean @@ -97,7 +97,7 @@ theorem smoothSheafCommRing.isUnit_stalk_iff {x : M} #adaptation_note /-- https://github.com/leanprover/lean4/pull/6024 was `exact`; somehow `convert` bypasess unification issues -/ convert ((contDiffAt_inv _ (hVf y)).contMDiffAt).comp y - (f.smooth.comp (smooth_inclusion hUV)).smoothAt + (f.contMDiff.comp (contMDiff_inclusion hUV)).contMDiffAt /-- The non-units of the stalk at `x` of the sheaf of smooth functions from `M` to `𝕜`, considered as a sheaf of commutative rings, are the functions whose values at `x` are zero. -/ diff --git a/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean b/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean index b25eee8e25950..0f528359115a1 100644 --- a/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean +++ b/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean @@ -126,7 +126,7 @@ def smoothSheaf.evalAt (x : TopCat.of M) (U : OpenNhds x) lemma smoothSheaf.eval_surjective (x : M) : Function.Surjective (smoothSheaf.eval IM I N x) := by apply TopCat.stalkToFiber_surjective intro n - exact ⟨⊤, fun _ ↦ n, smooth_const, rfl⟩ + exact ⟨⊤, fun _ ↦ n, contMDiff_const, rfl⟩ instance [Nontrivial N] (x : M) : Nontrivial ((smoothSheaf IM I M N).presheaf.stalk x) := (smoothSheaf.eval_surjective IM I N x).nontrivial @@ -138,11 +138,14 @@ variable {IM I N} smoothSheaf.eval IM I N (x : M) ((smoothSheaf IM I M N).presheaf.germ U x hx f) = f ⟨x, hx⟩ := TopCat.stalkToFiber_germ ((contDiffWithinAt_localInvariantProp ⊤).localPredicate M N) _ _ _ _ -lemma smoothSheaf.smooth_section {U : (Opens (TopCat.of M))ᵒᵖ} +lemma smoothSheaf.contMDiff_section {U : (Opens (TopCat.of M))ᵒᵖ} (f : (smoothSheaf IM I M N).presheaf.obj U) : - Smooth IM I f := + ContMDiff IM I ⊤ f := (contDiffWithinAt_localInvariantProp ⊤).section_spec _ _ _ _ +@[deprecated (since := "2024-11-21")] +alias smoothSheaf.smooth_section := smoothSheaf.contMDiff_section + end TypeCat section LieGroup @@ -213,7 +216,7 @@ noncomputable def smoothSheafCommGroup : TopCat.Sheaf CommGrp.{u} (TopCat.of M) @[to_additive "For a manifold `M` and a smooth homomorphism `φ` between abelian additive Lie groups `A`, `A'`, the 'left-composition-by-`φ`' morphism of sheaves from `smoothSheafAddCommGroup IM I M A` to `smoothSheafAddCommGroup IM I' M A'`."] -def smoothSheafCommGroup.compLeft (φ : A →* A') (hφ : Smooth I I' φ) : +def smoothSheafCommGroup.compLeft (φ : A →* A') (hφ : ContMDiff I I' ⊤ φ) : smoothSheafCommGroup IM I M A ⟶ smoothSheafCommGroup IM I' M A' := CategoryTheory.Sheaf.Hom.mk <| { app := fun _ ↦ CommGrp.ofHom <| SmoothMap.compLeftMonoidHom _ _ φ hφ diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean b/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean index dab21687a47a9..3fd4222c1daba 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean @@ -212,34 +212,29 @@ theorem contMDiff_proj : ContMDiff (IB.prod 𝓘(𝕜, F)) IB n (π F E) := fun rw [contMDiffAt_totalSpace] at this exact this.1 -theorem smooth_proj : Smooth (IB.prod 𝓘(𝕜, F)) IB (π F E) := - contMDiff_proj E +@[deprecated (since := "2024-11-21")] alias smooth_proj := contMDiff_proj theorem contMDiffOn_proj {s : Set (TotalSpace F E)} : ContMDiffOn (IB.prod 𝓘(𝕜, F)) IB n (π F E) s := (Bundle.contMDiff_proj E).contMDiffOn -theorem smoothOn_proj {s : Set (TotalSpace F E)} : SmoothOn (IB.prod 𝓘(𝕜, F)) IB (π F E) s := - contMDiffOn_proj E +@[deprecated (since := "2024-11-21")] alias smoothOn_proj := contMDiffOn_proj theorem contMDiffAt_proj {p : TotalSpace F E} : ContMDiffAt (IB.prod 𝓘(𝕜, F)) IB n (π F E) p := (Bundle.contMDiff_proj E).contMDiffAt -theorem smoothAt_proj {p : TotalSpace F E} : SmoothAt (IB.prod 𝓘(𝕜, F)) IB (π F E) p := - Bundle.contMDiffAt_proj E +@[deprecated (since := "2024-11-21")] alias smoothAt_proj := contMDiffAt_proj theorem contMDiffWithinAt_proj {s : Set (TotalSpace F E)} {p : TotalSpace F E} : ContMDiffWithinAt (IB.prod 𝓘(𝕜, F)) IB n (π F E) s p := (Bundle.contMDiffAt_proj E).contMDiffWithinAt -theorem smoothWithinAt_proj {s : Set (TotalSpace F E)} {p : TotalSpace F E} : - SmoothWithinAt (IB.prod 𝓘(𝕜, F)) IB (π F E) s p := - Bundle.contMDiffWithinAt_proj E +@[deprecated (since := "2024-11-21")] alias smoothWithinAt_proj := contMDiffWithinAt_proj variable (𝕜) [∀ x, AddCommMonoid (E x)] variable [∀ x, Module 𝕜 (E x)] [VectorBundle 𝕜 F E] -theorem smooth_zeroSection : Smooth IB (IB.prod 𝓘(𝕜, F)) (zeroSection F E) := fun x ↦ by +theorem contMDiff_zeroSection : ContMDiff IB (IB.prod 𝓘(𝕜, F)) ⊤ (zeroSection F E) := fun x ↦ by unfold zeroSection rw [Bundle.contMDiffAt_section] apply (contMDiffAt_const (c := 0)).congr_of_eventuallyEq @@ -247,6 +242,8 @@ theorem smooth_zeroSection : Smooth IB (IB.prod 𝓘(𝕜, F)) (zeroSection F E) (mem_baseSet_trivializationAt F E x)] with y hy using congr_arg Prod.snd <| (trivializationAt F E x).zeroSection 𝕜 hy +@[deprecated (since := "2024-11-21")] alias smooth_zeroSection := contMDiff_zeroSection + end Bundle end @@ -272,9 +269,9 @@ topological vector bundle over `B` with fibers isomorphic to `F`, then `SmoothVe registers that the bundle is smooth, in the sense of having smooth transition functions. This is a mixin, not carrying any new data. -/ class SmoothVectorBundle : Prop where - protected smoothOn_coordChangeL : + protected contMDiffOn_coordChangeL : ∀ (e e' : Trivialization F (π F E)) [MemTrivializationAtlas e] [MemTrivializationAtlas e'], - SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun b : B => (e.coordChangeL 𝕜 e' b : F →L[𝕜] F)) + ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun b : B => (e.coordChangeL 𝕜 e' b : F →L[𝕜] F)) (e.baseSet ∩ e'.baseSet) variable [SmoothVectorBundle F E IB] @@ -284,27 +281,23 @@ section SmoothCoordChange variable {F E} variable (e e' : Trivialization F (π F E)) [MemTrivializationAtlas e] [MemTrivializationAtlas e'] -theorem smoothOn_coordChangeL : - SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun b : B => (e.coordChangeL 𝕜 e' b : F →L[𝕜] F)) +theorem contMDiffOn_coordChangeL : + ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) n (fun b : B => (e.coordChangeL 𝕜 e' b : F →L[𝕜] F)) (e.baseSet ∩ e'.baseSet) := - SmoothVectorBundle.smoothOn_coordChangeL e e' + (SmoothVectorBundle.contMDiffOn_coordChangeL e e').of_le le_top -theorem smoothOn_symm_coordChangeL : - SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun b : B => ((e.coordChangeL 𝕜 e' b).symm : F →L[𝕜] F)) +theorem contMDiffOn_symm_coordChangeL : + ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) n (fun b : B => ((e.coordChangeL 𝕜 e' b).symm : F →L[𝕜] F)) (e.baseSet ∩ e'.baseSet) := by + apply ContMDiffOn.of_le _ le_top rw [inter_comm] - refine (SmoothVectorBundle.smoothOn_coordChangeL e' e).congr fun b hb ↦ ?_ + refine (SmoothVectorBundle.contMDiffOn_coordChangeL e' e).congr fun b hb ↦ ?_ rw [e.symm_coordChangeL e' hb] -theorem contMDiffOn_coordChangeL : - ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) n (fun b : B => (e.coordChangeL 𝕜 e' b : F →L[𝕜] F)) - (e.baseSet ∩ e'.baseSet) := - (smoothOn_coordChangeL e e').of_le le_top +@[deprecated (since := "2024-11-21")] alias smoothOn_coordChangeL := contMDiffOn_coordChangeL +@[deprecated (since := "2024-11-21")] +alias smoothOn_symm_coordChangeL := contMDiffOn_symm_coordChangeL -theorem contMDiffOn_symm_coordChangeL : - ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) n (fun b : B => ((e.coordChangeL 𝕜 e' b).symm : F →L[𝕜] F)) - (e.baseSet ∩ e'.baseSet) := - (smoothOn_symm_coordChangeL e e').of_le le_top variable {e e'} @@ -313,9 +306,7 @@ theorem contMDiffAt_coordChangeL {x : B} (h : x ∈ e.baseSet) (h' : x ∈ e'.ba (contMDiffOn_coordChangeL e e').contMDiffAt <| (e.open_baseSet.inter e'.open_baseSet).mem_nhds ⟨h, h'⟩ -theorem smoothAt_coordChangeL {x : B} (h : x ∈ e.baseSet) (h' : x ∈ e'.baseSet) : - SmoothAt IB 𝓘(𝕜, F →L[𝕜] F) (fun b : B => (e.coordChangeL 𝕜 e' b : F →L[𝕜] F)) x := - contMDiffAt_coordChangeL h h' +@[deprecated (since := "2024-11-21")] alias smoothAt_coordChangeL := contMDiffAt_coordChangeL variable {s : Set M} {f : M → B} {g : M → F} {x : M} @@ -339,25 +330,17 @@ protected theorem ContMDiff.coordChangeL ContMDiff IM 𝓘(𝕜, F →L[𝕜] F) n (fun y ↦ (e.coordChangeL 𝕜 e' (f y) : F →L[𝕜] F)) := fun x ↦ (hf x).coordChangeL (he x) (he' x) -protected nonrec theorem SmoothWithinAt.coordChangeL - (hf : SmoothWithinAt IM IB f s x) (he : f x ∈ e.baseSet) (he' : f x ∈ e'.baseSet) : - SmoothWithinAt IM 𝓘(𝕜, F →L[𝕜] F) (fun y ↦ (e.coordChangeL 𝕜 e' (f y) : F →L[𝕜] F)) s x := - hf.coordChangeL he he' +@[deprecated (since := "2024-11-21")] +alias SmoothWithinAt.coordChangeL := ContMDiffWithinAt.coordChangeL -protected nonrec theorem SmoothAt.coordChangeL - (hf : SmoothAt IM IB f x) (he : f x ∈ e.baseSet) (he' : f x ∈ e'.baseSet) : - SmoothAt IM 𝓘(𝕜, F →L[𝕜] F) (fun y ↦ (e.coordChangeL 𝕜 e' (f y) : F →L[𝕜] F)) x := - hf.coordChangeL he he' +@[deprecated (since := "2024-11-21")] +alias SmoothAt.coordChangeL := ContMDiffAt.coordChangeL -protected nonrec theorem SmoothOn.coordChangeL - (hf : SmoothOn IM IB f s) (he : MapsTo f s e.baseSet) (he' : MapsTo f s e'.baseSet) : - SmoothOn IM 𝓘(𝕜, F →L[𝕜] F) (fun y ↦ (e.coordChangeL 𝕜 e' (f y) : F →L[𝕜] F)) s := - hf.coordChangeL he he' +@[deprecated (since := "2024-11-21")] +alias SmoothOn.coordChangeL := ContMDiffOn.coordChangeL -protected nonrec theorem Smooth.coordChangeL - (hf : Smooth IM IB f) (he : ∀ x, f x ∈ e.baseSet) (he' : ∀ x, f x ∈ e'.baseSet) : - Smooth IM 𝓘(𝕜, F →L[𝕜] F) (fun y ↦ (e.coordChangeL 𝕜 e' (f y) : F →L[𝕜] F)) := - hf.coordChangeL he he' +@[deprecated (since := "2024-11-21")] +alias Smooth.coordChangeL := ContMDiff.coordChangeL protected theorem ContMDiffWithinAt.coordChange (hf : ContMDiffWithinAt IM IB n f s x) (hg : ContMDiffWithinAt IM 𝓘(𝕜, F) n g s x) @@ -386,26 +369,17 @@ protected theorem ContMDiff.coordChange (hf : ContMDiff IM IB n f) ContMDiff IM 𝓘(𝕜, F) n (fun y ↦ e.coordChange e' (f y) (g y)) := fun x ↦ (hf x).coordChange (hg x) (he x) (he' x) -protected nonrec theorem SmoothWithinAt.coordChange - (hf : SmoothWithinAt IM IB f s x) (hg : SmoothWithinAt IM 𝓘(𝕜, F) g s x) - (he : f x ∈ e.baseSet) (he' : f x ∈ e'.baseSet) : - SmoothWithinAt IM 𝓘(𝕜, F) (fun y ↦ e.coordChange e' (f y) (g y)) s x := - hf.coordChange hg he he' +@[deprecated (since := "2024-11-21")] +alias SmoothWithinAt.coordChange := ContMDiffWithinAt.coordChange -protected nonrec theorem SmoothAt.coordChange (hf : SmoothAt IM IB f x) - (hg : SmoothAt IM 𝓘(𝕜, F) g x) (he : f x ∈ e.baseSet) (he' : f x ∈ e'.baseSet) : - SmoothAt IM 𝓘(𝕜, F) (fun y ↦ e.coordChange e' (f y) (g y)) x := - hf.coordChange hg he he' +@[deprecated (since := "2024-11-21")] +alias SmoothAt.coordChange := ContMDiffAt.coordChange -protected nonrec theorem SmoothOn.coordChange (hf : SmoothOn IM IB f s) - (hg : SmoothOn IM 𝓘(𝕜, F) g s) (he : MapsTo f s e.baseSet) (he' : MapsTo f s e'.baseSet) : - SmoothOn IM 𝓘(𝕜, F) (fun y ↦ e.coordChange e' (f y) (g y)) s := - hf.coordChange hg he he' +@[deprecated (since := "2024-11-21")] +alias SmoothOn.coordChange := ContMDiffOn.coordChange -protected theorem Smooth.coordChange (hf : Smooth IM IB f) - (hg : Smooth IM 𝓘(𝕜, F) g) (he : ∀ x, f x ∈ e.baseSet) (he' : ∀ x, f x ∈ e'.baseSet) : - Smooth IM 𝓘(𝕜, F) (fun y ↦ e.coordChange e' (f y) (g y)) := fun x ↦ - (hf x).coordChange (hg x) (he x) (he' x) +@[deprecated (since := "2024-11-21")] +alias Smooth.coordChange := ContMDiff.coordChange variable (e e') @@ -458,8 +432,8 @@ instance SmoothFiberwiseLinear.hasGroupoid : haveI : MemTrivializationAtlas e := ⟨he⟩ haveI : MemTrivializationAtlas e' := ⟨he'⟩ rw [mem_smoothFiberwiseLinear_iff] - refine ⟨_, _, e.open_baseSet.inter e'.open_baseSet, smoothOn_coordChangeL e e', - smoothOn_symm_coordChangeL e e', ?_⟩ + refine ⟨_, _, e.open_baseSet.inter e'.open_baseSet, contMDiffOn_coordChangeL e e', + contMDiffOn_symm_coordChangeL e e', ?_⟩ refine PartialHomeomorph.eqOnSourceSetoid.symm ⟨?_, ?_⟩ · simp only [e.symm_trans_source_eq e', FiberwiseLinear.partialHomeomorph, trans_toPartialEquiv, symm_toPartialEquiv] @@ -478,10 +452,10 @@ instance Bundle.TotalSpace.smoothManifoldWithCorners [SmoothManifoldWithCorners refine ⟨ContMDiffOn.congr ?_ (EqOnSource.eqOn heφ), ContMDiffOn.congr ?_ (EqOnSource.eqOn (EqOnSource.symm' heφ))⟩ · rw [EqOnSource.source_eq heφ] - apply smoothOn_fst.prod_mk + apply contMDiffOn_fst.prod_mk exact (hφ.comp contMDiffOn_fst <| prod_subset_preimage_fst _ _).clm_apply contMDiffOn_snd · rw [EqOnSource.target_eq heφ] - apply smoothOn_fst.prod_mk + apply contMDiffOn_fst.prod_mk exact (h2φ.comp contMDiffOn_fst <| prod_subset_preimage_fst _ _).clm_apply contMDiffOn_snd section @@ -517,41 +491,36 @@ theorem Trivialization.contMDiff_iff {f : M → TotalSpace F E} (he : ∀ x, f x ContMDiff IM 𝓘(𝕜, F) n (fun x ↦ (e (f x)).2) := (forall_congr' fun x ↦ e.contMDiffAt_iff (he x)).trans forall_and -theorem Trivialization.smoothWithinAt_iff {f : M → TotalSpace F E} {s : Set M} {x₀ : M} - (he : f x₀ ∈ e.source) : - SmoothWithinAt IM (IB.prod 𝓘(𝕜, F)) f s x₀ ↔ - SmoothWithinAt IM IB (fun x => (f x).proj) s x₀ ∧ - SmoothWithinAt IM 𝓘(𝕜, F) (fun x ↦ (e (f x)).2) s x₀ := - e.contMDiffWithinAt_iff he +@[deprecated (since := "2024-11-21")] +alias Trivialization.smoothWithinAt_iff := Trivialization.contMDiffWithinAt_iff -theorem Trivialization.smoothAt_iff {f : M → TotalSpace F E} {x₀ : M} (he : f x₀ ∈ e.source) : - SmoothAt IM (IB.prod 𝓘(𝕜, F)) f x₀ ↔ - SmoothAt IM IB (fun x => (f x).proj) x₀ ∧ SmoothAt IM 𝓘(𝕜, F) (fun x ↦ (e (f x)).2) x₀ := - e.contMDiffAt_iff he +@[deprecated (since := "2024-11-21")] +alias Trivialization.smoothAt_iff := Trivialization.contMDiffAt_iff -theorem Trivialization.smoothOn_iff {f : M → TotalSpace F E} {s : Set M} - (he : MapsTo f s e.source) : - SmoothOn IM (IB.prod 𝓘(𝕜, F)) f s ↔ - SmoothOn IM IB (fun x => (f x).proj) s ∧ SmoothOn IM 𝓘(𝕜, F) (fun x ↦ (e (f x)).2) s := - e.contMDiffOn_iff he - -theorem Trivialization.smooth_iff {f : M → TotalSpace F E} (he : ∀ x, f x ∈ e.source) : - Smooth IM (IB.prod 𝓘(𝕜, F)) f ↔ - Smooth IM IB (fun x => (f x).proj) ∧ Smooth IM 𝓘(𝕜, F) (fun x ↦ (e (f x)).2) := - e.contMDiff_iff he - -theorem Trivialization.smoothOn (e : Trivialization F (π F E)) [MemTrivializationAtlas e] : - SmoothOn (IB.prod 𝓘(𝕜, F)) (IB.prod 𝓘(𝕜, F)) e e.source := by - have : SmoothOn (IB.prod 𝓘(𝕜, F)) (IB.prod 𝓘(𝕜, F)) id e.source := smoothOn_id - rw [e.smoothOn_iff (mapsTo_id _)] at this +@[deprecated (since := "2024-11-21")] +alias Trivialization.smoothOn_iff := Trivialization.contMDiffOn_iff + +@[deprecated (since := "2024-11-21")] +alias Trivialization.smooth_iff := Trivialization.contMDiff_iff + +theorem Trivialization.contMDiffOn (e : Trivialization F (π F E)) [MemTrivializationAtlas e] : + ContMDiffOn (IB.prod 𝓘(𝕜, F)) (IB.prod 𝓘(𝕜, F)) ⊤ e e.source := by + have : ContMDiffOn (IB.prod 𝓘(𝕜, F)) (IB.prod 𝓘(𝕜, F)) ⊤ id e.source := contMDiffOn_id + rw [e.contMDiffOn_iff (mapsTo_id _)] at this exact (this.1.prod_mk this.2).congr fun x hx ↦ (e.mk_proj_snd hx).symm -theorem Trivialization.smoothOn_symm (e : Trivialization F (π F E)) [MemTrivializationAtlas e] : - SmoothOn (IB.prod 𝓘(𝕜, F)) (IB.prod 𝓘(𝕜, F)) e.toPartialHomeomorph.symm e.target := by - rw [e.smoothOn_iff e.toPartialHomeomorph.symm_mapsTo] - refine ⟨smoothOn_fst.congr fun x hx ↦ e.proj_symm_apply hx, smoothOn_snd.congr fun x hx ↦ ?_⟩ +theorem Trivialization.contMDiffOn_symm (e : Trivialization F (π F E)) [MemTrivializationAtlas e] : + ContMDiffOn (IB.prod 𝓘(𝕜, F)) (IB.prod 𝓘(𝕜, F)) ⊤ e.toPartialHomeomorph.symm e.target := by + rw [e.contMDiffOn_iff e.toPartialHomeomorph.symm_mapsTo] + refine ⟨contMDiffOn_fst.congr fun x hx ↦ e.proj_symm_apply hx, + contMDiffOn_snd.congr fun x hx ↦ ?_⟩ rw [e.apply_symm_apply hx] +@[deprecated (since := "2024-11-21")] alias Trivialization.smoothOn := Trivialization.contMDiffOn + +@[deprecated (since := "2024-11-21")] +alias Trivialization.smoothOn_symm := Trivialization.contMDiffOn_symm + end /-! ### Core construction for smooth vector bundles -/ @@ -563,21 +532,24 @@ variable {ι : Type*} (Z : VectorBundleCore 𝕜 B F ι) /-- Mixin for a `VectorBundleCore` stating smoothness (of transition functions). -/ class IsSmooth (IB : ModelWithCorners 𝕜 EB HB) : Prop where - smoothOn_coordChange : - ∀ i j, SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (Z.coordChange i j) (Z.baseSet i ∩ Z.baseSet j) + contMDiffOn_coordChange : + ∀ i j, ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (Z.coordChange i j) (Z.baseSet i ∩ Z.baseSet j) -theorem smoothOn_coordChange (IB : ModelWithCorners 𝕜 EB HB) [h : Z.IsSmooth IB] (i j : ι) : - SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (Z.coordChange i j) (Z.baseSet i ∩ Z.baseSet j) := +theorem contMDiffOn_coordChange (IB : ModelWithCorners 𝕜 EB HB) [h : Z.IsSmooth IB] (i j : ι) : + ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (Z.coordChange i j) (Z.baseSet i ∩ Z.baseSet j) := h.1 i j +@[deprecated (since := "2024-11-21")] +alias smoothOn_coordChange := contMDiffOn_coordChange + variable [Z.IsSmooth IB] /-- If a `VectorBundleCore` has the `IsSmooth` mixin, then the vector bundle constructed from it is a smooth vector bundle. -/ instance smoothVectorBundle : SmoothVectorBundle F Z.Fiber IB where - smoothOn_coordChangeL := by + contMDiffOn_coordChangeL := by rintro - - ⟨i, rfl⟩ ⟨i', rfl⟩ - refine (Z.smoothOn_coordChange IB i i').congr fun b hb ↦ ?_ + refine (Z.contMDiffOn_coordChange IB i i').congr fun b hb ↦ ?_ ext v exact Z.localTriv_coordChange_eq i i' hb v @@ -587,12 +559,12 @@ end VectorBundleCore /-- A trivial vector bundle over a smooth manifold is a smooth vector bundle. -/ instance Bundle.Trivial.smoothVectorBundle : SmoothVectorBundle F (Bundle.Trivial B F) IB where - smoothOn_coordChangeL := by + contMDiffOn_coordChangeL := by intro e e' he he' obtain rfl := Bundle.Trivial.eq_trivialization B F e obtain rfl := Bundle.Trivial.eq_trivialization B F e' simp_rw [Bundle.Trivial.trivialization.coordChangeL] - exact smooth_const.smoothOn + exact contMDiff_const.contMDiffOn /-! ### Direct sums of smooth vector bundles -/ @@ -613,15 +585,14 @@ variable [SmoothManifoldWithCorners IB B] /-- The direct sum of two smooth vector bundles over the same base is a smooth vector bundle. -/ instance Bundle.Prod.smoothVectorBundle : SmoothVectorBundle (F₁ × F₂) (E₁ ×ᵇ E₂) IB where - smoothOn_coordChangeL := by + contMDiffOn_coordChangeL := by rintro _ _ ⟨e₁, e₂, i₁, i₂, rfl⟩ ⟨e₁', e₂', i₁', i₂', rfl⟩ - rw [SmoothOn] refine ContMDiffOn.congr ?_ (e₁.coordChangeL_prod 𝕜 e₁' e₂ e₂') refine ContMDiffOn.clm_prodMap ?_ ?_ - · refine (smoothOn_coordChangeL e₁ e₁').mono ?_ + · refine (contMDiffOn_coordChangeL e₁ e₁').mono ?_ simp only [Trivialization.baseSet_prod, mfld_simps] mfld_set_tac - · refine (smoothOn_coordChangeL e₂ e₂').mono ?_ + · refine (contMDiffOn_coordChangeL e₂ e₂').mono ?_ simp only [Trivialization.baseSet_prod, mfld_simps] mfld_set_tac @@ -641,7 +612,7 @@ class IsSmooth (a : VectorPrebundle 𝕜 F E) : Prop where exists_smoothCoordChange : ∀ᵉ (e ∈ a.pretrivializationAtlas) (e' ∈ a.pretrivializationAtlas), ∃ f : B → F →L[𝕜] F, - SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) f (e.baseSet ∩ e'.baseSet) ∧ + ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ f (e.baseSet ∩ e'.baseSet) ∧ ∀ (b : B) (_ : b ∈ e.baseSet ∩ e'.baseSet) (v : F), f b v = (e' ⟨b, e.symm b v⟩).2 @@ -655,11 +626,14 @@ noncomputable def smoothCoordChange (he : e ∈ a.pretrivializationAtlas) (he' : e' ∈ a.pretrivializationAtlas) (b : B) : F →L[𝕜] F := Classical.choose (ha.exists_smoothCoordChange e he e' he') b -theorem smoothOn_smoothCoordChange (he : e ∈ a.pretrivializationAtlas) +theorem contMDiffOn_smoothCoordChange (he : e ∈ a.pretrivializationAtlas) (he' : e' ∈ a.pretrivializationAtlas) : - SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (a.smoothCoordChange IB he he') (e.baseSet ∩ e'.baseSet) := + ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (a.smoothCoordChange IB he he') (e.baseSet ∩ e'.baseSet) := (Classical.choose_spec (ha.exists_smoothCoordChange e he e' he')).1 +@[deprecated (since := "2024-11-21")] +alias smoothOn_smoothCoordChange := contMDiffOn_smoothCoordChange + theorem smoothCoordChange_apply (he : e ∈ a.pretrivializationAtlas) (he' : e' ∈ a.pretrivializationAtlas) {b : B} (hb : b ∈ e.baseSet ∩ e'.baseSet) (v : F) : a.smoothCoordChange IB he he' b v = (e' ⟨b, e.symm b v⟩).2 := @@ -678,9 +652,9 @@ variable (IB) in theorem smoothVectorBundle : @SmoothVectorBundle _ _ F E _ _ _ _ _ _ IB _ _ _ _ _ _ a.totalSpaceTopology _ a.toFiberBundle a.toVectorBundle := letI := a.totalSpaceTopology; letI := a.toFiberBundle; letI := a.toVectorBundle - { smoothOn_coordChangeL := by + { contMDiffOn_coordChangeL := by rintro _ _ ⟨e, he, rfl⟩ ⟨e', he', rfl⟩ - refine (a.smoothOn_smoothCoordChange he he').congr ?_ + refine (a.contMDiffOn_smoothCoordChange he he').congr ?_ intro b hb ext v rw [a.smoothCoordChange_apply he he' hb v, ContinuousLinearEquiv.coe_coe, diff --git a/Mathlib/Geometry/Manifold/VectorBundle/FiberwiseLinear.lean b/Mathlib/Geometry/Manifold/VectorBundle/FiberwiseLinear.lean index 2a0e077c860aa..0bc3d2285bba8 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/FiberwiseLinear.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/FiberwiseLinear.lean @@ -103,14 +103,14 @@ fiberwise linear partial homeomorphism. -/ theorem SmoothFiberwiseLinear.locality_aux₁ (e : PartialHomeomorph (B × F) (B × F)) (h : ∀ p ∈ e.source, ∃ s : Set (B × F), IsOpen s ∧ p ∈ s ∧ ∃ (φ : B → F ≃L[𝕜] F) (u : Set B) (hu : IsOpen u) - (hφ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => (φ x : F →L[𝕜] F)) u) - (h2φ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => ((φ x).symm : F →L[𝕜] F)) u), + (hφ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => (φ x : F →L[𝕜] F)) u) + (h2φ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => ((φ x).symm : F →L[𝕜] F)) u), (e.restr s).EqOnSource (FiberwiseLinear.partialHomeomorph φ hu hφ.continuousOn h2φ.continuousOn)) : ∃ U : Set B, e.source = U ×ˢ univ ∧ ∀ x ∈ U, ∃ (φ : B → F ≃L[𝕜] F) (u : Set B) (hu : IsOpen u) (_huU : u ⊆ U) (_hux : x ∈ u), - ∃ (hφ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => (φ x : F →L[𝕜] F)) u) - (h2φ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => ((φ x).symm : F →L[𝕜] F)) u), + ∃ (hφ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => (φ x : F →L[𝕜] F)) u) + (h2φ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => ((φ x).symm : F →L[𝕜] F)) u), (e.restr (u ×ˢ univ)).EqOnSource (FiberwiseLinear.partialHomeomorph φ hu hφ.continuousOn h2φ.continuousOn) := by rw [SetCoe.forall'] at h @@ -155,13 +155,13 @@ theorem SmoothFiberwiseLinear.locality_aux₂ (e : PartialHomeomorph (B × F) (B (hU : e.source = U ×ˢ univ) (h : ∀ x ∈ U, ∃ (φ : B → F ≃L[𝕜] F) (u : Set B) (hu : IsOpen u) (_hUu : u ⊆ U) (_hux : x ∈ u) - (hφ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => (φ x : F →L[𝕜] F)) u) - (h2φ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => ((φ x).symm : F →L[𝕜] F)) u), + (hφ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => (φ x : F →L[𝕜] F)) u) + (h2φ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => ((φ x).symm : F →L[𝕜] F)) u), (e.restr (u ×ˢ univ)).EqOnSource (FiberwiseLinear.partialHomeomorph φ hu hφ.continuousOn h2φ.continuousOn)) : ∃ (Φ : B → F ≃L[𝕜] F) (U : Set B) (hU₀ : IsOpen U) (hΦ : - SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => (Φ x : F →L[𝕜] F)) U) (h2Φ : - SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => ((Φ x).symm : F →L[𝕜] F)) U), + ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => (Φ x : F →L[𝕜] F)) U) (h2Φ : + ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => ((Φ x).symm : F →L[𝕜] F)) U), e.EqOnSource (FiberwiseLinear.partialHomeomorph Φ hU₀ hΦ.continuousOn h2Φ.continuousOn) := by classical rw [SetCoe.forall'] at h @@ -192,14 +192,14 @@ theorem SmoothFiberwiseLinear.locality_aux₂ (e : PartialHomeomorph (B × F) (B intro x y hyu refine (hΦ y (hUu x hyu)).trans ?_ exact iUnionLift_mk ⟨y, hyu⟩ _ - have hΦ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun y => (Φ y : F →L[𝕜] F)) U := by + have hΦ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun y => (Φ y : F →L[𝕜] F)) U := by apply contMDiffOn_of_locally_contMDiffOn intro x hx refine ⟨u ⟨x, hx⟩, hu ⟨x, hx⟩, hux _, ?_⟩ refine (ContMDiffOn.congr (hφ ⟨x, hx⟩) ?_).mono inter_subset_right intro y hy rw [hΦφ ⟨x, hx⟩ y hy] - have h2Φ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun y => ((Φ y).symm : F →L[𝕜] F)) U := by + have h2Φ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun y => ((Φ y).symm : F →L[𝕜] F)) U := by apply contMDiffOn_of_locally_contMDiffOn intro x hx refine ⟨u ⟨x, hx⟩, hu ⟨x, hx⟩, hux _, ?_⟩ @@ -221,13 +221,13 @@ variable {F B IB} in -- in `smoothFiberwiseLinear` are quite slow, even with this change) private theorem mem_aux {e : PartialHomeomorph (B × F) (B × F)} : (e ∈ ⋃ (φ : B → F ≃L[𝕜] F) (U : Set B) (hU : IsOpen U) - (hφ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => φ x : B → F →L[𝕜] F) U) - (h2φ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => (φ x).symm : B → F →L[𝕜] F) U), + (hφ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => φ x : B → F →L[𝕜] F) U) + (h2φ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => (φ x).symm : B → F →L[𝕜] F) U), {e | e.EqOnSource (FiberwiseLinear.partialHomeomorph φ hU hφ.continuousOn h2φ.continuousOn)}) ↔ ∃ (φ : B → F ≃L[𝕜] F) (U : Set B) (hU : IsOpen U) - (hφ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => φ x : B → F →L[𝕜] F) U) - (h2φ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => (φ x).symm : B → F →L[𝕜] F) U), + (hφ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => φ x : B → F →L[𝕜] F) U) + (h2φ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => (φ x).symm : B → F →L[𝕜] F) U), e.EqOnSource (FiberwiseLinear.partialHomeomorph φ hU hφ.continuousOn h2φ.continuousOn) := by simp only [mem_iUnion, mem_setOf_eq] @@ -239,8 +239,8 @@ to the vector bundle belong to this groupoid. -/ def smoothFiberwiseLinear : StructureGroupoid (B × F) where members := ⋃ (φ : B → F ≃L[𝕜] F) (U : Set B) (hU : IsOpen U) - (hφ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => φ x : B → F →L[𝕜] F) U) - (h2φ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => (φ x).symm : B → F →L[𝕜] F) U), + (hφ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => φ x : B → F →L[𝕜] F) U) + (h2φ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => (φ x).symm : B → F →L[𝕜] F) U), {e | e.EqOnSource (FiberwiseLinear.partialHomeomorph φ hU hφ.continuousOn h2φ.continuousOn)} trans' := by simp only [mem_aux] @@ -248,11 +248,11 @@ def smoothFiberwiseLinear : StructureGroupoid (B × F) where refine ⟨fun b => (φ b).trans (φ' b), _, hU.inter hU', ?_, ?_, Setoid.trans (PartialHomeomorph.EqOnSource.trans' heφ heφ') ⟨?_, ?_⟩⟩ · show - SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) + ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x : B => (φ' x).toContinuousLinearMap ∘L (φ x).toContinuousLinearMap) (U ∩ U') exact (hφ'.mono inter_subset_right).clm_comp (hφ.mono inter_subset_left) · show - SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) + ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x : B => (φ x).symm.toContinuousLinearMap ∘L (φ' x).symm.toContinuousLinearMap) (U ∩ U') exact (h2φ.mono inter_subset_left).clm_comp (h2φ'.mono inter_subset_right) @@ -267,8 +267,8 @@ def smoothFiberwiseLinear : StructureGroupoid (B × F) where exact hφ id_mem' := by simp_rw [mem_aux] - refine ⟨fun _ ↦ ContinuousLinearEquiv.refl 𝕜 F, univ, isOpen_univ, smoothOn_const, - smoothOn_const, ⟨?_, fun b _hb ↦ rfl⟩⟩ + refine ⟨fun _ ↦ ContinuousLinearEquiv.refl 𝕜 F, univ, isOpen_univ, contMDiffOn_const, + contMDiffOn_const, ⟨?_, fun b _hb ↦ rfl⟩⟩ simp only [FiberwiseLinear.partialHomeomorph, PartialHomeomorph.refl_partialEquiv, PartialEquiv.refl_source, univ_prod_univ] locality' := by @@ -286,7 +286,7 @@ def smoothFiberwiseLinear : StructureGroupoid (B × F) where theorem mem_smoothFiberwiseLinear_iff (e : PartialHomeomorph (B × F) (B × F)) : e ∈ smoothFiberwiseLinear B F IB ↔ ∃ (φ : B → F ≃L[𝕜] F) (U : Set B) (hU : IsOpen U) (hφ : - SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => φ x : B → F →L[𝕜] F) U) (h2φ : - SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => (φ x).symm : B → F →L[𝕜] F) U), + ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => φ x : B → F →L[𝕜] F) U) (h2φ : + ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => (φ x).symm : B → F →L[𝕜] F) U), e.EqOnSource (FiberwiseLinear.partialHomeomorph φ hU hφ.continuousOn h2φ.continuousOn) := mem_aux diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Hom.lean b/Mathlib/Geometry/Manifold/VectorBundle/Hom.lean index f56284f714605..d00cefedfa42b 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Hom.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Hom.lean @@ -41,16 +41,19 @@ variable {𝕜 B F₁ F₂ M : Type*} {E₁ : B → Type*} {E₂ : B → Type*} local notation "LE₁E₂" => TotalSpace (F₁ →L[𝕜] F₂) (Bundle.ContinuousLinearMap (RingHom.id 𝕜) E₁ E₂) -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11083): moved slow parts to separate lemmas -theorem smoothOn_continuousLinearMapCoordChange +theorem contMDiffOn_continuousLinearMapCoordChange [SmoothVectorBundle F₁ E₁ IB] [SmoothVectorBundle F₂ E₂ IB] [MemTrivializationAtlas e₁] [MemTrivializationAtlas e₁'] [MemTrivializationAtlas e₂] [MemTrivializationAtlas e₂'] : - SmoothOn IB 𝓘(𝕜, (F₁ →L[𝕜] F₂) →L[𝕜] F₁ →L[𝕜] F₂) + ContMDiffOn IB 𝓘(𝕜, (F₁ →L[𝕜] F₂) →L[𝕜] F₁ →L[𝕜] F₂) ⊤ (continuousLinearMapCoordChange (RingHom.id 𝕜) e₁ e₁' e₂ e₂') (e₁.baseSet ∩ e₂.baseSet ∩ (e₁'.baseSet ∩ e₂'.baseSet)) := by - have h₁ := smoothOn_coordChangeL (IB := IB) e₁' e₁ - have h₂ := smoothOn_coordChangeL (IB := IB) e₂ e₂' + have h₁ := contMDiffOn_coordChangeL (IB := IB) e₁' e₁ (n := ⊤) + have h₂ := contMDiffOn_coordChangeL (IB := IB) e₂ e₂' (n := ⊤) refine (h₁.mono ?_).cle_arrowCongr (h₂.mono ?_) <;> mfld_set_tac +@[deprecated (since := "2024-11-21")] +alias smoothOn_continuousLinearMapCoordChange := contMDiffOn_continuousLinearMapCoordChange + variable [∀ x, TopologicalAddGroup (E₂ x)] [∀ x, ContinuousSMul 𝕜 (E₂ x)] theorem hom_chart (y₀ y : LE₁E₂) : @@ -67,12 +70,8 @@ theorem contMDiffAt_hom_bundle (f : M → LE₁E₂) {x₀ : M} {n : ℕ∞} : (fun x => inCoordinates F₁ E₁ F₂ E₂ (f x₀).1 (f x).1 (f x₀).1 (f x).1 (f x).2) x₀ := contMDiffAt_totalSpace .. -theorem smoothAt_hom_bundle (f : M → LE₁E₂) {x₀ : M} : - SmoothAt IM (IB.prod 𝓘(𝕜, F₁ →L[𝕜] F₂)) f x₀ ↔ - SmoothAt IM IB (fun x => (f x).1) x₀ ∧ - SmoothAt IM 𝓘(𝕜, F₁ →L[𝕜] F₂) - (fun x => inCoordinates F₁ E₁ F₂ E₂ (f x₀).1 (f x).1 (f x₀).1 (f x).1 (f x).2) x₀ := - contMDiffAt_hom_bundle f +@[deprecated (since := "2024-11-21")] alias smoothAt_hom_bundle := contMDiffAt_hom_bundle + variable [SmoothVectorBundle F₁ E₁ IB] [SmoothVectorBundle F₂ E₂ IB] @@ -81,7 +80,7 @@ instance Bundle.ContinuousLinearMap.vectorPrebundle.isSmooth : exists_smoothCoordChange := by rintro _ ⟨e₁, e₂, he₁, he₂, rfl⟩ _ ⟨e₁', e₂', he₁', he₂', rfl⟩ exact ⟨continuousLinearMapCoordChange (RingHom.id 𝕜) e₁ e₁' e₂ e₂', - smoothOn_continuousLinearMapCoordChange, + contMDiffOn_continuousLinearMapCoordChange, continuousLinearMapCoordChange_apply (RingHom.id 𝕜) e₁ e₁' e₂ e₂'⟩ instance SmoothVectorBundle.continuousLinearMap : diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Pullback.lean b/Mathlib/Geometry/Manifold/VectorBundle/Pullback.lean index ef7514b05054c..7ba8cdb7b86d3 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Pullback.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Pullback.lean @@ -28,14 +28,14 @@ variable [NontriviallyNormedField 𝕜] [∀ x, AddCommMonoid (E x)] [∀ x, Mod [ChartedSpace HB B] {EB' : Type*} [NormedAddCommGroup EB'] [NormedSpace 𝕜 EB'] {HB' : Type*} [TopologicalSpace HB'] (IB' : ModelWithCorners 𝕜 EB' HB') [TopologicalSpace B'] [ChartedSpace HB' B'] [FiberBundle F E] - [VectorBundle 𝕜 F E] [SmoothVectorBundle F E IB] (f : SmoothMap IB' IB B' B) + [VectorBundle 𝕜 F E] [SmoothVectorBundle F E IB] (f : ContMDiffMap IB' IB B' B ⊤) /-- For a smooth vector bundle `E` over a manifold `B` and a smooth map `f : B' → B`, the pullback vector bundle `f *ᵖ E` is a smooth vector bundle. -/ instance SmoothVectorBundle.pullback : SmoothVectorBundle F (f *ᵖ E) IB' where - smoothOn_coordChangeL := by + contMDiffOn_coordChangeL := by rintro _ _ ⟨e, he, rfl⟩ ⟨e', he', rfl⟩ - refine ((smoothOn_coordChangeL e e').comp f.smooth.smoothOn fun b hb => hb).congr ?_ + refine ((contMDiffOn_coordChangeL e e').comp f.contMDiff.contMDiffOn fun b hb => hb).congr ?_ rintro b (hb : f b ∈ e.baseSet ∩ e'.baseSet); ext v show ((e.pullback f).coordChangeL 𝕜 (e'.pullback f) b) v = (e.coordChangeL 𝕜 e' (f b)) v rw [e.coordChangeL_apply e' hb, (e.pullback f).coordChangeL_apply' _] diff --git a/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean b/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean index 198e5170ad73e..751bf15036342 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean @@ -39,9 +39,7 @@ structure ContMDiffSection where protected contMDiff_toFun : ContMDiff I (I.prod 𝓘(𝕜, F)) n fun x ↦ TotalSpace.mk' F x (toFun x) -/-- Bundled smooth sections of a vector bundle. -/ -abbrev SmoothSection := - ContMDiffSection I F ⊤ V +@[deprecated (since := "024-11-21")] alias SmoothSection := ContMDiffSection @[inherit_doc] scoped[Manifold] notation "Cₛ^" n "⟮" I "; " F ", " V "⟯" => ContMDiffSection I F n V @@ -65,9 +63,7 @@ protected theorem contMDiff (s : Cₛ^n⟮I; F, V⟯) : ContMDiff I (I.prod 𝓘(𝕜, F)) n fun x => TotalSpace.mk' F x (s x : V x) := s.contMDiff_toFun -protected theorem smooth (s : Cₛ^∞⟮I; F, V⟯) : - Smooth I (I.prod 𝓘(𝕜, F)) fun x => TotalSpace.mk' F x (s x : V x) := - s.contMDiff_toFun +@[deprecated (since := "2024-11-21")] alias smooth := ContMDiffSection.contMDiff theorem coe_inj ⦃s t : Cₛ^n⟮I; F, V⟯⦄ (h : (s : ∀ x, V x) = t) : s = t := DFunLike.ext' h @@ -114,7 +110,7 @@ theorem coe_sub (s t : Cₛ^n⟮I; F, V⟯) : ⇑(s - t) = s - t := rfl instance instZero : Zero Cₛ^n⟮I; F, V⟯ := - ⟨⟨fun _ => 0, (smooth_zeroSection 𝕜 V).of_le le_top⟩⟩ + ⟨⟨fun _ => 0, (contMDiff_zeroSection 𝕜 V).of_le le_top⟩⟩ instance inhabited : Inhabited Cₛ^n⟮I; F, V⟯ := ⟨0⟩ diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean index fd5bda4a24a80..f40f1f950c07a 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean @@ -286,7 +286,7 @@ end TangentBundle instance tangentBundleCore.isSmooth : (tangentBundleCore I M).IsSmooth I := by refine ⟨fun i j => ?_⟩ - rw [SmoothOn, contMDiffOn_iff_source_of_mem_maximalAtlas (subset_maximalAtlas i.2), + rw [contMDiffOn_iff_source_of_mem_maximalAtlas (subset_maximalAtlas i.2), contMDiffOn_iff_contDiffOn] · refine ((contDiffOn_fderiv_coord_change (I := I) i j).congr fun x hx => ?_).mono ?_ · rw [PartialEquiv.trans_source'] at hx diff --git a/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean b/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean index 2c5558360b39f..5551413ea3416 100644 --- a/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean +++ b/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean @@ -52,7 +52,7 @@ def embeddingPiTangent : C^∞⟮I, M; 𝓘(ℝ, ι → E × ℝ), ι → E × val x i := (f i x • extChartAt I (f.c i) x, f i x) property := contMDiff_pi_space.2 fun i => - ((f i).smooth_smul contMDiffOn_extChartAt).prod_mk_space (f i).smooth + ((f i).contMDiff_smul contMDiffOn_extChartAt).prod_mk_space (f i).contMDiff @[local simp] theorem embeddingPiTangent_coe : @@ -81,7 +81,8 @@ theorem comp_embeddingPiTangent_mfderiv (x : M) (hx : x ∈ s) : set L := (ContinuousLinearMap.fst ℝ E ℝ).comp (@ContinuousLinearMap.proj ℝ _ ι (fun _ => E × ℝ) _ _ (fun _ => inferInstance) (f.ind x hx)) - have := L.hasMFDerivAt.comp x f.embeddingPiTangent.smooth.mdifferentiableAt.hasMFDerivAt + have := L.hasMFDerivAt.comp x + (f.embeddingPiTangent.contMDiff.mdifferentiableAt le_top).hasMFDerivAt convert hasMFDerivAt_unique this _ refine (hasMFDerivAt_extChartAt (f.mem_chartAt_ind_source x hx)).congr_of_eventuallyEq ?_ refine (f.eventuallyEq_one x hx).mono fun y hy => ?_ @@ -106,7 +107,7 @@ supports of bump functions, then for some `n` it can be immersed into the `n`-di Euclidean space. -/ theorem exists_immersion_euclidean {ι : Type*} [Finite ι] (f : SmoothBumpCovering ι I M) : ∃ (n : ℕ) (e : M → EuclideanSpace ℝ (Fin n)), - Smooth I (𝓡 n) e ∧ Injective e ∧ ∀ x : M, Injective (mfderiv I (𝓡 n) e x) := by + ContMDiff I (𝓡 n) ⊤ e ∧ Injective e ∧ ∀ x : M, Injective (mfderiv I (𝓡 n) e x) := by cases nonempty_fintype ι set F := EuclideanSpace ℝ (Fin <| finrank ℝ (ι → E × ℝ)) letI : IsNoetherian ℝ (E × ℝ) := IsNoetherian.iff_fg.2 inferInstance @@ -114,10 +115,10 @@ theorem exists_immersion_euclidean {ι : Type*} [Finite ι] (f : SmoothBumpCover set eEF : (ι → E × ℝ) ≃L[ℝ] F := ContinuousLinearEquiv.ofFinrankEq finrank_euclideanSpace_fin.symm refine ⟨_, eEF ∘ f.embeddingPiTangent, - eEF.toDiffeomorph.smooth.comp f.embeddingPiTangent.smooth, + eEF.toDiffeomorph.contMDiff.comp f.embeddingPiTangent.contMDiff, eEF.injective.comp f.embeddingPiTangent_injective, fun x => ?_⟩ rw [mfderiv_comp _ eEF.differentiableAt.mdifferentiableAt - f.embeddingPiTangent.smooth.mdifferentiableAt, + (f.embeddingPiTangent.contMDiff.mdifferentiableAt le_top), eEF.mfderiv_eq] exact eEF.injective.comp (f.embeddingPiTangent_injective_mfderiv _ trivial) @@ -128,7 +129,7 @@ supports of bump functions, then for some `n` it can be embedded into the `n`-di Euclidean space. -/ theorem exists_embedding_euclidean_of_compact [T2Space M] [CompactSpace M] : ∃ (n : ℕ) (e : M → EuclideanSpace ℝ (Fin n)), - Smooth I (𝓡 n) e ∧ IsClosedEmbedding e ∧ ∀ x : M, Injective (mfderiv I (𝓡 n) e x) := by + ContMDiff I (𝓡 n) ⊤ e ∧ IsClosedEmbedding e ∧ ∀ x : M, Injective (mfderiv I (𝓡 n) e x) := by rcases SmoothBumpCovering.exists_isSubordinate I isClosed_univ fun (x : M) _ => univ_mem with ⟨ι, f, -⟩ haveI := f.fintype From f801de4934457721044d87bd482f67a3219c30c1 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 25 Nov 2024 09:33:04 +0000 Subject: [PATCH 274/829] chore: split notation typeclasses out of Algebra.Group.Defs (#19350) --- Mathlib.lean | 1 + Mathlib/Algebra/Group/Defs.lean | 192 +---------------------- Mathlib/Algebra/Group/Operations.lean | 215 ++++++++++++++++++++++++++ Mathlib/Data/Part.lean | 2 +- 4 files changed, 219 insertions(+), 191 deletions(-) create mode 100644 Mathlib/Algebra/Group/Operations.lean diff --git a/Mathlib.lean b/Mathlib.lean index 56e4a03916d8e..4b515136d2aa3 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -287,6 +287,7 @@ import Mathlib.Algebra.Group.Nat.Even import Mathlib.Algebra.Group.Nat.TypeTags import Mathlib.Algebra.Group.Nat.Units import Mathlib.Algebra.Group.NatPowAssoc +import Mathlib.Algebra.Group.Operations import Mathlib.Algebra.Group.Opposite import Mathlib.Algebra.Group.PNatPowAssoc import Mathlib.Algebra.Group.Pi.Basic diff --git a/Mathlib/Algebra/Group/Defs.lean b/Mathlib/Algebra/Group/Defs.lean index 3abfe30a7d417..3d069aab663a3 100644 --- a/Mathlib/Algebra/Group/Defs.lean +++ b/Mathlib/Algebra/Group/Defs.lean @@ -6,12 +6,9 @@ Authors: Jeremy Avigad, Leonardo de Moura, Simon Hudon, Mario Carneiro import Mathlib.Data.Int.Notation import Mathlib.Data.Nat.BinaryRec import Mathlib.Algebra.Group.ZeroOne +import Mathlib.Algebra.Group.Operations import Mathlib.Logic.Function.Defs -import Mathlib.Tactic.Lemma -import Mathlib.Tactic.TypeStar import Mathlib.Tactic.Simps.Basic -import Mathlib.Tactic.ToAdditive -import Mathlib.Util.AssertExists import Batteries.Logic /-! @@ -29,21 +26,15 @@ The file does not contain any lemmas except for For basic lemmas about these classes see `Algebra.Group.Basic`. -We also introduce notation classes `SMul` and `VAdd` for multiplicative and additive -actions and register the following instances: +We register the following instances: - `Pow M ℕ`, for monoids `M`, and `Pow G ℤ` for groups `G`; - `SMul ℕ M` for additive monoids `M`, and `SMul ℤ G` for additive groups `G`. -`SMul` is typically, but not exclusively, used for scalar multiplication-like operators. -See the module `Algebra.AddTorsor` for a motivating example for the name `VAdd` (vector addition). - ## Notation - `+`, `-`, `*`, `/`, `^` : the usual arithmetic operations; the underlying functions are `Add.add`, `Neg.neg`/`Sub.sub`, `Mul.mul`, `Div.div`, and `HPow.hPow`. -- `a • b` is used as notation for `HSMul.hSMul a b`. -- `a +ᵥ b` is used as notation for `HVAdd.hVAdd a b`. -/ @@ -55,187 +46,8 @@ universe u v w open Function -/-- -The notation typeclass for heterogeneous additive actions. -This enables the notation `a +ᵥ b : γ` where `a : α`, `b : β`. --/ -class HVAdd (α : Type u) (β : Type v) (γ : outParam (Type w)) where - /-- `a +ᵥ b` computes the sum of `a` and `b`. - The meaning of this notation is type-dependent. -/ - hVAdd : α → β → γ - -/-- -The notation typeclass for heterogeneous scalar multiplication. -This enables the notation `a • b : γ` where `a : α`, `b : β`. - -It is assumed to represent a left action in some sense. -The notation `a • b` is augmented with a macro (below) to have it elaborate as a left action. -Only the `b` argument participates in the elaboration algorithm: the algorithm uses the type of `b` -when calculating the type of the surrounding arithmetic expression -and it tries to insert coercions into `b` to get some `b'` -such that `a • b'` has the same type as `b'`. -See the module documentation near the macro for more details. --/ -class HSMul (α : Type u) (β : Type v) (γ : outParam (Type w)) where - /-- `a • b` computes the product of `a` and `b`. - The meaning of this notation is type-dependent, but it is intended to be used for left actions. -/ - hSMul : α → β → γ - -attribute [notation_class smul Simps.copySecond] HSMul -attribute [notation_class nsmul Simps.nsmulArgs] HSMul -attribute [notation_class zsmul Simps.zsmulArgs] HSMul - -/-- Type class for the `+ᵥ` notation. -/ -class VAdd (G : Type u) (P : Type v) where - /-- `a +ᵥ b` computes the sum of `a` and `b`. The meaning of this notation is type-dependent, - but it is intended to be used for left actions. -/ - vadd : G → P → P - -/-- Type class for the `-ᵥ` notation. -/ -class VSub (G : outParam Type*) (P : Type*) where - /-- `a -ᵥ b` computes the difference of `a` and `b`. The meaning of this notation is - type-dependent, but it is intended to be used for additive torsors. -/ - vsub : P → P → G - -/-- Typeclass for types with a scalar multiplication operation, denoted `•` (`\bu`) -/ -@[to_additive (attr := ext)] -class SMul (M : Type u) (α : Type v) where - /-- `a • b` computes the product of `a` and `b`. The meaning of this notation is type-dependent, - but it is intended to be used for left actions. -/ - smul : M → α → α - -@[inherit_doc] infixl:65 " +ᵥ " => HVAdd.hVAdd -@[inherit_doc] infixl:65 " -ᵥ " => VSub.vsub -@[inherit_doc] infixr:73 " • " => HSMul.hSMul - -/-! -We have a macro to make `x • y` notation participate in the expression tree elaborator, -like other arithmetic expressions such as `+`, `*`, `/`, `^`, `=`, inequalities, etc. -The macro is using the `leftact%` elaborator introduced in -[this RFC](https://github.com/leanprover/lean4/issues/2854). - -As a concrete example of the effect of this macro, consider -```lean -variable [Ring R] [AddCommMonoid M] [Module R M] (r : R) (N : Submodule R M) (m : M) (n : N) -#check m + r • n -``` -Without the macro, the expression would elaborate as `m + ↑(r • n : ↑N) : M`. -With the macro, the expression elaborates as `m + r • (↑n : M) : M`. -To get the first interpretation, one can write `m + (r • n :)`. - -Here is a quick review of the expression tree elaborator: -1. It builds up an expression tree of all the immediately accessible operations - that are marked with `binop%`, `unop%`, `leftact%`, `rightact%`, `binrel%`, etc. -2. It elaborates every leaf term of this tree - (without an expected type, so as if it were temporarily wrapped in `(... :)`). -3. Using the types of each elaborated leaf, it computes a supremum type they can all be - coerced to, if such a supremum exists. -4. It inserts coercions around leaf terms wherever needed. - -The hypothesis is that individual expression trees tend to be calculations with respect -to a single algebraic structure. - -Note(kmill): If we were to remove `HSMul` and switch to using `SMul` directly, -then the expression tree elaborator would not be able to insert coercions within the right operand; -they would likely appear as `↑(x • y)` rather than `x • ↑y`, unlike other arithmetic operations. --/ - -@[inherit_doc HSMul.hSMul] -macro_rules | `($x • $y) => `(leftact% HSMul.hSMul $x $y) - -attribute [to_additive existing] Mul Div HMul instHMul HDiv instHDiv HSMul -attribute [to_additive (reorder := 1 2) SMul] Pow -attribute [to_additive (reorder := 1 2)] HPow -attribute [to_additive existing (reorder := 1 2, 5 6) hSMul] HPow.hPow -attribute [to_additive existing (reorder := 1 2, 4 5) smul] Pow.pow - -@[to_additive (attr := default_instance)] -instance instHSMul {α β} [SMul α β] : HSMul α β β where - hSMul := SMul.smul - -@[to_additive] -theorem SMul.smul_eq_hSMul {α β} [SMul α β] : (SMul.smul : α → β → β) = HSMul.hSMul := rfl - -attribute [to_additive existing (reorder := 1 2)] instHPow - variable {G : Type*} -/-- Class of types that have an inversion operation. -/ -@[to_additive, notation_class] -class Inv (α : Type u) where - /-- Invert an element of α. -/ - inv : α → α - -@[inherit_doc] -postfix:max "⁻¹" => Inv.inv - -section ite -variable {α : Type*} (P : Prop) [Decidable P] - -section Mul -variable [Mul α] - -@[to_additive] -lemma mul_dite (a : α) (b : P → α) (c : ¬ P → α) : - (a * if h : P then b h else c h) = if h : P then a * b h else a * c h := by split <;> rfl - -@[to_additive] -lemma mul_ite (a b c : α) : (a * if P then b else c) = if P then a * b else a * c := mul_dite .. - -@[to_additive] -lemma dite_mul (a : P → α) (b : ¬ P → α) (c : α) : - (if h : P then a h else b h) * c = if h : P then a h * c else b h * c := by split <;> rfl - -@[to_additive] -lemma ite_mul (a b c : α) : (if P then a else b) * c = if P then a * c else b * c := dite_mul .. - --- We make `mul_ite` and `ite_mul` simp lemmas, but not `add_ite` or `ite_add`. --- The problem we're trying to avoid is dealing with sums of the form `∑ x ∈ s, (f x + ite P 1 0)`, --- in which `add_ite` followed by `sum_ite` would needlessly slice up --- the `f x` terms according to whether `P` holds at `x`. --- There doesn't appear to be a corresponding difficulty so far with `mul_ite` and `ite_mul`. -attribute [simp] mul_dite dite_mul mul_ite ite_mul - -@[to_additive] -lemma dite_mul_dite (a : P → α) (b : ¬ P → α) (c : P → α) (d : ¬ P → α) : - ((if h : P then a h else b h) * if h : P then c h else d h) = - if h : P then a h * c h else b h * d h := by split <;> rfl - -@[to_additive] -lemma ite_mul_ite (a b c d : α) : - ((if P then a else b) * if P then c else d) = if P then a * c else b * d := by split <;> rfl - -end Mul - -section Div -variable [Div α] - -@[to_additive] -lemma div_dite (a : α) (b : P → α) (c : ¬ P → α) : - (a / if h : P then b h else c h) = if h : P then a / b h else a / c h := by split <;> rfl - -@[to_additive] -lemma div_ite (a b c : α) : (a / if P then b else c) = if P then a / b else a / c := div_dite .. - -@[to_additive] -lemma dite_div (a : P → α) (b : ¬ P → α) (c : α) : - (if h : P then a h else b h) / c = if h : P then a h / c else b h / c := by split <;> rfl - -@[to_additive] -lemma ite_div (a b c : α) : (if P then a else b) / c = if P then a / c else b / c := dite_div .. - -@[to_additive] -lemma dite_div_dite (a : P → α) (b : ¬ P → α) (c : P → α) (d : ¬ P → α) : - ((if h : P then a h else b h) / if h : P then c h else d h) = - if h : P then a h / c h else b h / d h := by split <;> rfl - -@[to_additive] -lemma ite_div_ite (a b c d : α) : - ((if P then a else b) / if P then c else d) = if P then a / c else b / d := dite_div_dite .. - -end Div -end ite - section Mul variable [Mul G] diff --git a/Mathlib/Algebra/Group/Operations.lean b/Mathlib/Algebra/Group/Operations.lean new file mode 100644 index 0000000000000..dcc6957a258a7 --- /dev/null +++ b/Mathlib/Algebra/Group/Operations.lean @@ -0,0 +1,215 @@ +/- +Copyright (c) 2014 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura, Simon Hudon, Mario Carneiro +-/ + +import Mathlib.Tactic.Lemma +import Mathlib.Tactic.TypeStar +import Mathlib.Tactic.ToAdditive +import Mathlib.Util.AssertExists + +/-! +# Typeclasses for algebraic operations + +Notation typeclass for `Inv`, the multiplicative analogue of `Neg`. + +We also introduce notation classes `SMul` and `VAdd` for multiplicative and additive +actions. + +`SMul` is typically, but not exclusively, used for scalar multiplication-like operators. +See the module `Algebra.AddTorsor` for a motivating example for the name `VAdd` (vector addition). + +## Notation + +- `a • b` is used as notation for `HSMul.hSMul a b`. +- `a +ᵥ b` is used as notation for `HVAdd.hVAdd a b`. + +-/ + +assert_not_exists One +assert_not_exists Function.Injective + +universe u v w + + +/-- +The notation typeclass for heterogeneous additive actions. +This enables the notation `a +ᵥ b : γ` where `a : α`, `b : β`. +-/ +class HVAdd (α : Type u) (β : Type v) (γ : outParam (Type w)) where + /-- `a +ᵥ b` computes the sum of `a` and `b`. + The meaning of this notation is type-dependent. -/ + hVAdd : α → β → γ + +/-- +The notation typeclass for heterogeneous scalar multiplication. +This enables the notation `a • b : γ` where `a : α`, `b : β`. + +It is assumed to represent a left action in some sense. +The notation `a • b` is augmented with a macro (below) to have it elaborate as a left action. +Only the `b` argument participates in the elaboration algorithm: the algorithm uses the type of `b` +when calculating the type of the surrounding arithmetic expression +and it tries to insert coercions into `b` to get some `b'` +such that `a • b'` has the same type as `b'`. +See the module documentation near the macro for more details. +-/ +class HSMul (α : Type u) (β : Type v) (γ : outParam (Type w)) where + /-- `a • b` computes the product of `a` and `b`. + The meaning of this notation is type-dependent, but it is intended to be used for left actions. -/ + hSMul : α → β → γ + +attribute [notation_class smul Simps.copySecond] HSMul +attribute [notation_class nsmul Simps.nsmulArgs] HSMul +attribute [notation_class zsmul Simps.zsmulArgs] HSMul + +/-- Type class for the `+ᵥ` notation. -/ +class VAdd (G : Type u) (P : Type v) where + /-- `a +ᵥ b` computes the sum of `a` and `b`. The meaning of this notation is type-dependent, + but it is intended to be used for left actions. -/ + vadd : G → P → P + +/-- Type class for the `-ᵥ` notation. -/ +class VSub (G : outParam Type*) (P : Type*) where + /-- `a -ᵥ b` computes the difference of `a` and `b`. The meaning of this notation is + type-dependent, but it is intended to be used for additive torsors. -/ + vsub : P → P → G + +/-- Typeclass for types with a scalar multiplication operation, denoted `•` (`\bu`) -/ +@[to_additive (attr := ext)] +class SMul (M : Type u) (α : Type v) where + /-- `a • b` computes the product of `a` and `b`. The meaning of this notation is type-dependent, + but it is intended to be used for left actions. -/ + smul : M → α → α + +@[inherit_doc] infixl:65 " +ᵥ " => HVAdd.hVAdd +@[inherit_doc] infixl:65 " -ᵥ " => VSub.vsub +@[inherit_doc] infixr:73 " • " => HSMul.hSMul + +/-! +We have a macro to make `x • y` notation participate in the expression tree elaborator, +like other arithmetic expressions such as `+`, `*`, `/`, `^`, `=`, inequalities, etc. +The macro is using the `leftact%` elaborator introduced in +[this RFC](https://github.com/leanprover/lean4/issues/2854). + +As a concrete example of the effect of this macro, consider +```lean +variable [Ring R] [AddCommMonoid M] [Module R M] (r : R) (N : Submodule R M) (m : M) (n : N) +#check m + r • n +``` +Without the macro, the expression would elaborate as `m + ↑(r • n : ↑N) : M`. +With the macro, the expression elaborates as `m + r • (↑n : M) : M`. +To get the first interpretation, one can write `m + (r • n :)`. + +Here is a quick review of the expression tree elaborator: +1. It builds up an expression tree of all the immediately accessible operations + that are marked with `binop%`, `unop%`, `leftact%`, `rightact%`, `binrel%`, etc. +2. It elaborates every leaf term of this tree + (without an expected type, so as if it were temporarily wrapped in `(... :)`). +3. Using the types of each elaborated leaf, it computes a supremum type they can all be + coerced to, if such a supremum exists. +4. It inserts coercions around leaf terms wherever needed. + +The hypothesis is that individual expression trees tend to be calculations with respect +to a single algebraic structure. + +Note(kmill): If we were to remove `HSMul` and switch to using `SMul` directly, +then the expression tree elaborator would not be able to insert coercions within the right operand; +they would likely appear as `↑(x • y)` rather than `x • ↑y`, unlike other arithmetic operations. +-/ + +@[inherit_doc HSMul.hSMul] +macro_rules | `($x • $y) => `(leftact% HSMul.hSMul $x $y) + +attribute [to_additive existing] Mul Div HMul instHMul HDiv instHDiv HSMul +attribute [to_additive (reorder := 1 2) SMul] Pow +attribute [to_additive (reorder := 1 2)] HPow +attribute [to_additive existing (reorder := 1 2, 5 6) hSMul] HPow.hPow +attribute [to_additive existing (reorder := 1 2, 4 5) smul] Pow.pow + +@[to_additive (attr := default_instance)] +instance instHSMul {α β} [SMul α β] : HSMul α β β where + hSMul := SMul.smul + +@[to_additive] +theorem SMul.smul_eq_hSMul {α β} [SMul α β] : (SMul.smul : α → β → β) = HSMul.hSMul := rfl + +attribute [to_additive existing (reorder := 1 2)] instHPow + +variable {G : Type*} + +/-- Class of types that have an inversion operation. -/ +@[to_additive, notation_class] +class Inv (α : Type u) where + /-- Invert an element of α. -/ + inv : α → α + +@[inherit_doc] +postfix:max "⁻¹" => Inv.inv + +section ite +variable {α : Type*} (P : Prop) [Decidable P] + +section Mul +variable [Mul α] + +@[to_additive] +lemma mul_dite (a : α) (b : P → α) (c : ¬ P → α) : + (a * if h : P then b h else c h) = if h : P then a * b h else a * c h := by split <;> rfl + +@[to_additive] +lemma mul_ite (a b c : α) : (a * if P then b else c) = if P then a * b else a * c := mul_dite .. + +@[to_additive] +lemma dite_mul (a : P → α) (b : ¬ P → α) (c : α) : + (if h : P then a h else b h) * c = if h : P then a h * c else b h * c := by split <;> rfl + +@[to_additive] +lemma ite_mul (a b c : α) : (if P then a else b) * c = if P then a * c else b * c := dite_mul .. + +-- We make `mul_ite` and `ite_mul` simp lemmas, but not `add_ite` or `ite_add`. +-- The problem we're trying to avoid is dealing with sums of the form `∑ x ∈ s, (f x + ite P 1 0)`, +-- in which `add_ite` followed by `sum_ite` would needlessly slice up +-- the `f x` terms according to whether `P` holds at `x`. +-- There doesn't appear to be a corresponding difficulty so far with `mul_ite` and `ite_mul`. +attribute [simp] mul_dite dite_mul mul_ite ite_mul + +@[to_additive] +lemma dite_mul_dite (a : P → α) (b : ¬ P → α) (c : P → α) (d : ¬ P → α) : + ((if h : P then a h else b h) * if h : P then c h else d h) = + if h : P then a h * c h else b h * d h := by split <;> rfl + +@[to_additive] +lemma ite_mul_ite (a b c d : α) : + ((if P then a else b) * if P then c else d) = if P then a * c else b * d := by split <;> rfl + +end Mul + +section Div +variable [Div α] + +@[to_additive] +lemma div_dite (a : α) (b : P → α) (c : ¬ P → α) : + (a / if h : P then b h else c h) = if h : P then a / b h else a / c h := by split <;> rfl + +@[to_additive] +lemma div_ite (a b c : α) : (a / if P then b else c) = if P then a / b else a / c := div_dite .. + +@[to_additive] +lemma dite_div (a : P → α) (b : ¬ P → α) (c : α) : + (if h : P then a h else b h) / c = if h : P then a h / c else b h / c := by split <;> rfl + +@[to_additive] +lemma ite_div (a b c : α) : (if P then a else b) / c = if P then a / c else b / c := dite_div .. + +@[to_additive] +lemma dite_div_dite (a : P → α) (b : ¬ P → α) (c : P → α) (d : ¬ P → α) : + ((if h : P then a h else b h) / if h : P then c h else d h) = + if h : P then a h / c h else b h / d h := by split <;> rfl + +@[to_additive] +lemma ite_div_ite (a b c d : α) : + ((if P then a else b) / if P then c else d) = if P then a / c else b / d := dite_div_dite .. + +end Div +end ite diff --git a/Mathlib/Data/Part.lean b/Mathlib/Data/Part.lean index 58395794466e5..842a941412357 100644 --- a/Mathlib/Data/Part.lean +++ b/Mathlib/Data/Part.lean @@ -5,7 +5,7 @@ Authors: Mario Carneiro, Jeremy Avigad, Simon Hudon -/ import Mathlib.Data.Set.Subsingleton import Mathlib.Logic.Equiv.Defs -import Mathlib.Algebra.Group.Defs +import Mathlib.Algebra.Group.Operations /-! # Partial values of a type From 725d2ded25e22f4f26c9554e47a48398db571c9b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Nov 2024 10:13:39 +0000 Subject: [PATCH 275/829] chore: rename `CompleteLattice.Independent`/`.SetIndependent` to `iSupIndep`/`sSupIndep` (#19409) These names are shorter, closer to `SupIndep` and clearer without their namespace. --- Counterexamples/DirectSumIsInternal.lean | 8 +- Mathlib/Algebra/DirectSum/Basic.lean | 4 +- Mathlib/Algebra/DirectSum/Internal.lean | 2 +- Mathlib/Algebra/DirectSum/LinearMap.lean | 12 +- Mathlib/Algebra/DirectSum/Module.lean | 73 +++--- Mathlib/Algebra/DirectSum/Ring.lean | 5 +- Mathlib/Algebra/Lie/Semisimple/Basic.lean | 6 +- Mathlib/Algebra/Lie/Semisimple/Defs.lean | 2 +- Mathlib/Algebra/Lie/Submodule.lean | 9 +- Mathlib/Algebra/Lie/TraceForm.lean | 8 +- Mathlib/Algebra/Lie/Weights/Basic.lean | 33 ++- Mathlib/Algebra/Lie/Weights/Chain.lean | 6 +- Mathlib/Algebra/Lie/Weights/Killing.lean | 4 +- Mathlib/Algebra/Module/Torsion.lean | 15 +- Mathlib/Analysis/InnerProductSpace/Basic.lean | 4 +- .../InnerProductSpace/Projection.lean | 2 +- Mathlib/GroupTheory/NoncommPiCoprod.lean | 28 ++- Mathlib/GroupTheory/Sylow.lean | 2 +- Mathlib/LinearAlgebra/DFinsupp.lean | 110 +++++---- Mathlib/LinearAlgebra/Dimension/Finite.lean | 21 +- Mathlib/LinearAlgebra/Eigenspace/Basic.lean | 16 +- Mathlib/LinearAlgebra/Eigenspace/Matrix.lean | 2 +- Mathlib/LinearAlgebra/Eigenspace/Pi.lean | 4 +- Mathlib/LinearAlgebra/LinearIndependent.lean | 11 +- .../Projectivization/Independence.lean | 8 +- Mathlib/Order/CompactlyGenerated/Basic.lean | 127 ++++++---- Mathlib/Order/SupIndep.lean | 223 ++++++++++++------ Mathlib/RingTheory/FiniteLength.lean | 8 +- Mathlib/RingTheory/Noetherian/Defs.lean | 13 +- Mathlib/RingTheory/SimpleModule.lean | 10 +- scripts/nolints_prime_decls.txt | 12 +- 31 files changed, 486 insertions(+), 302 deletions(-) diff --git a/Counterexamples/DirectSumIsInternal.lean b/Counterexamples/DirectSumIsInternal.lean index db6389188d856..7632fb6414a31 100644 --- a/Counterexamples/DirectSumIsInternal.lean +++ b/Counterexamples/DirectSumIsInternal.lean @@ -12,10 +12,10 @@ import Mathlib.Tactic.FinCases # Not all complementary decompositions of a module over a semiring make up a direct sum This shows that while `ℤ≤0` and `ℤ≥0` are complementary `ℕ`-submodules of `ℤ`, which in turn -implies as a collection they are `CompleteLattice.Independent` and that they span all of `ℤ`, they +implies as a collection they are `iSupIndep` and that they span all of `ℤ`, they do not form a decomposition into a direct sum. -This file demonstrates why `DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top` must +This file demonstrates why `DirectSum.isInternal_submodule_of_iSupIndep_of_iSup_eq_top` must take `Ring R` and not `Semiring R`. -/ @@ -57,9 +57,9 @@ theorem withSign.isCompl : IsCompl ℤ≥0 ℤ≤0 := by · exact Submodule.mem_sup_left (mem_withSign_one.mpr hp) · exact Submodule.mem_sup_right (mem_withSign_neg_one.mpr hn) -def withSign.independent : CompleteLattice.Independent withSign := by +def withSign.independent : iSupIndep withSign := by apply - (CompleteLattice.independent_pair UnitsInt.one_ne_neg_one _).mpr withSign.isCompl.disjoint + (iSupIndep_pair UnitsInt.one_ne_neg_one _).mpr withSign.isCompl.disjoint intro i fin_cases i <;> simp diff --git a/Mathlib/Algebra/DirectSum/Basic.lean b/Mathlib/Algebra/DirectSum/Basic.lean index 4aab3891bf560..50045a6c1ba5b 100644 --- a/Mathlib/Algebra/DirectSum/Basic.lean +++ b/Mathlib/Algebra/DirectSum/Basic.lean @@ -367,8 +367,8 @@ theorem coe_of_apply {M S : Type*} [DecidableEq ι] [AddCommMonoid M] [SetLike S `M` is said to be internal if the canonical map `(⨁ i, A i) →+ M` is bijective. For the alternate statement in terms of independence and spanning, see -`DirectSum.subgroup_isInternal_iff_independent_and_supr_eq_top` and -`DirectSum.isInternal_submodule_iff_independent_and_iSup_eq_top`. -/ +`DirectSum.subgroup_isInternal_iff_iSupIndep_and_supr_eq_top` and +`DirectSum.isInternal_submodule_iff_iSupIndep_and_iSup_eq_top`. -/ def IsInternal {M S : Type*} [DecidableEq ι] [AddCommMonoid M] [SetLike S M] [AddSubmonoidClass S M] (A : ι → S) : Prop := Function.Bijective (DirectSum.coeAddMonoidHom A) diff --git a/Mathlib/Algebra/DirectSum/Internal.lean b/Mathlib/Algebra/DirectSum/Internal.lean index 34fbe957a6174..a75722faa88dc 100644 --- a/Mathlib/Algebra/DirectSum/Internal.lean +++ b/Mathlib/Algebra/DirectSum/Internal.lean @@ -28,7 +28,7 @@ to represent this case, `(h : DirectSum.IsInternal A) [SetLike.GradedMonoid A]` needed. In the future there will likely be a data-carrying, constructive, typeclass version of `DirectSum.IsInternal` for providing an explicit decomposition function. -When `CompleteLattice.Independent (Set.range A)` (a weaker condition than +When `iSupIndep (Set.range A)` (a weaker condition than `DirectSum.IsInternal A`), these provide a grading of `⨆ i, A i`, and the mapping `⨁ i, A i →+ ⨆ i, A i` can be obtained as `DirectSum.toAddMonoid (fun i ↦ AddSubmonoid.inclusion <| le_iSup A i)`. diff --git a/Mathlib/Algebra/DirectSum/LinearMap.lean b/Mathlib/Algebra/DirectSum/LinearMap.lean index 25ffc82193de4..ff8f8b3d653a4 100644 --- a/Mathlib/Algebra/DirectSum/LinearMap.lean +++ b/Mathlib/Algebra/DirectSum/LinearMap.lean @@ -82,8 +82,8 @@ lemma trace_eq_zero_of_mapsTo_ne (h : IsInternal N) [IsNoetherian R M] (σ : ι → ι) (hσ : ∀ i, σ i ≠ i) {f : Module.End R M} (hf : ∀ i, MapsTo f (N i) (N <| σ i)) : trace R M f = 0 := by - have hN : {i | N i ≠ ⊥}.Finite := CompleteLattice.WellFoundedGT.finite_ne_bot_of_independent - h.submodule_independent + have hN : {i | N i ≠ ⊥}.Finite := WellFoundedGT.finite_ne_bot_of_iSupIndep + h.submodule_iSupIndep let s := hN.toFinset let κ := fun i ↦ Module.Free.ChooseBasisIndex R (N i) let b : (i : s) → Basis (κ i) R (N i) := fun i ↦ Module.Free.chooseBasis R (N i) @@ -109,10 +109,10 @@ lemma trace_comp_eq_zero_of_commute_of_trace_restrict_eq_zero (f.mapsTo_maxGenEigenspace_of_comm rfl μ) suffices ∀ μ, trace R _ ((g ∘ₗ f).restrict (hfg μ)) = 0 by classical - have hds := DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top + have hds := DirectSum.isInternal_submodule_of_iSupIndep_of_iSup_eq_top f.independent_maxGenEigenspace hf have h_fin : {μ | f.maxGenEigenspace μ ≠ ⊥}.Finite := - CompleteLattice.WellFoundedGT.finite_ne_bot_of_independent f.independent_maxGenEigenspace + WellFoundedGT.finite_ne_bot_of_iSupIndep f.independent_maxGenEigenspace simp [trace_eq_sum_trace_restrict' hds h_fin hfg, this] intro μ replace h_comm : Commute (g.restrict (f.mapsTo_maxGenEigenspace_of_comm h_comm μ)) @@ -136,14 +136,14 @@ Note that it is important the statement gives the user definitional control over _type_ of the term `trace R p (f.restrict hp')` depends on `p`. -/ lemma trace_eq_sum_trace_restrict_of_eq_biSup [∀ i, Module.Finite R (N i)] [∀ i, Module.Free R (N i)] - (s : Finset ι) (h : CompleteLattice.Independent <| fun i : s ↦ N i) + (s : Finset ι) (h : iSupIndep <| fun i : s ↦ N i) {f : Module.End R M} (hf : ∀ i, MapsTo f (N i) (N i)) (p : Submodule R M) (hp : p = ⨆ i ∈ s, N i) (hp' : MapsTo f p p := hp ▸ mapsTo_biSup_of_mapsTo (s : Set ι) hf) : trace R p (f.restrict hp') = ∑ i ∈ s, trace R (N i) (f.restrict (hf i)) := by classical let N' : s → Submodule R p := fun i ↦ (N i).comap p.subtype - replace h : IsInternal N' := hp ▸ isInternal_biSup_submodule_of_independent (s : Set ι) h + replace h : IsInternal N' := hp ▸ isInternal_biSup_submodule_of_iSupIndep (s : Set ι) h have hf' : ∀ i, MapsTo (restrict f hp') (N' i) (N' i) := fun i x hx' ↦ by simpa using hf i hx' let e : (i : s) → N' i ≃ₗ[R] N i := fun ⟨i, hi⟩ ↦ (N i).comapSubtypeEquivOfLe (hp ▸ le_biSup N hi) have _i1 : ∀ i, Module.Finite R (N' i) := fun i ↦ Module.Finite.equiv (e i).symm diff --git a/Mathlib/Algebra/DirectSum/Module.lean b/Mathlib/Algebra/DirectSum/Module.lean index 299c5d4caad54..dba8c5158d5d0 100644 --- a/Mathlib/Algebra/DirectSum/Module.lean +++ b/Mathlib/Algebra/DirectSum/Module.lean @@ -319,8 +319,11 @@ theorem IsInternal.submodule_iSup_eq_top (h : IsInternal A) : iSup A = ⊤ := by exact Function.Bijective.surjective h /-- If a direct sum of submodules is internal then the submodules are independent. -/ -theorem IsInternal.submodule_independent (h : IsInternal A) : CompleteLattice.Independent A := - CompleteLattice.independent_of_dfinsupp_lsum_injective _ h.injective +theorem IsInternal.submodule_iSupIndep (h : IsInternal A) : iSupIndep A := + iSupIndep_of_dfinsupp_lsum_injective _ h.injective + +@[deprecated (since := "2024-11-24")] +alias IsInternal.submodule_independent := IsInternal.submodule_iSupIndep /-- Given an internal direct sum decomposition of a module `M`, and a basis for each of the components of the direct sum, the disjoint union of these bases is a basis for `M`. -/ @@ -374,7 +377,7 @@ the two submodules are complementary. Over a `Ring R`, this is true as an iff, a `DirectSum.isInternal_submodule_iff_isCompl`. -/ theorem IsInternal.isCompl {A : ι → Submodule R M} {i j : ι} (hij : i ≠ j) (h : (Set.univ : Set ι) = {i, j}) (hi : IsInternal A) : IsCompl (A i) (A j) := - ⟨hi.submodule_independent.pairwiseDisjoint hij, + ⟨hi.submodule_iSupIndep.pairwiseDisjoint hij, codisjoint_iff.mpr <| Eq.symm <| hi.submodule_iSup_eq_top.symm.trans <| by rw [← sSup_pair, iSup, ← Set.image_univ, h, Set.image_insert_eq, Set.image_singleton]⟩ @@ -387,60 +390,76 @@ variable {ι : Type v} [dec_ι : DecidableEq ι] variable {M : Type*} [AddCommGroup M] [Module R M] /-- Note that this is not generally true for `[Semiring R]`; see -`CompleteLattice.Independent.dfinsupp_lsum_injective` for details. -/ -theorem isInternal_submodule_of_independent_of_iSup_eq_top {A : ι → Submodule R M} - (hi : CompleteLattice.Independent A) (hs : iSup A = ⊤) : IsInternal A := +`iSupIndep.dfinsupp_lsum_injective` for details. -/ +theorem isInternal_submodule_of_iSupIndep_of_iSup_eq_top {A : ι → Submodule R M} + (hi : iSupIndep A) (hs : iSup A = ⊤) : IsInternal A := ⟨hi.dfinsupp_lsum_injective, -- Note: https://github.com/leanprover-community/mathlib4/pull/8386 had to specify value of `f` (LinearMap.range_eq_top (f := DFinsupp.lsum _ _)).1 <| (Submodule.iSup_eq_range_dfinsupp_lsum _).symm.trans hs⟩ -/-- `iff` version of `DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top`, -`DirectSum.IsInternal.submodule_independent`, and `DirectSum.IsInternal.submodule_iSup_eq_top`. -/ -theorem isInternal_submodule_iff_independent_and_iSup_eq_top (A : ι → Submodule R M) : - IsInternal A ↔ CompleteLattice.Independent A ∧ iSup A = ⊤ := - ⟨fun i ↦ ⟨i.submodule_independent, i.submodule_iSup_eq_top⟩, - And.rec isInternal_submodule_of_independent_of_iSup_eq_top⟩ +@[deprecated (since := "2024-11-24")] +alias isInternal_submodule_of_independent_of_iSup_eq_top := + isInternal_submodule_of_iSupIndep_of_iSup_eq_top + +/-- `iff` version of `DirectSum.isInternal_submodule_of_iSupIndep_of_iSup_eq_top`, +`DirectSum.IsInternal.iSupIndep`, and `DirectSum.IsInternal.submodule_iSup_eq_top`. -/ +theorem isInternal_submodule_iff_iSupIndep_and_iSup_eq_top (A : ι → Submodule R M) : + IsInternal A ↔ iSupIndep A ∧ iSup A = ⊤ := + ⟨fun i ↦ ⟨i.submodule_iSupIndep, i.submodule_iSup_eq_top⟩, + And.rec isInternal_submodule_of_iSupIndep_of_iSup_eq_top⟩ + +@[deprecated (since := "2024-11-24")] +alias isInternal_submodule_iff_independent_and_iSup_eq_top := + isInternal_submodule_iff_iSupIndep_and_iSup_eq_top /-- If a collection of submodules has just two indices, `i` and `j`, then `DirectSum.IsInternal` is equivalent to `isCompl`. -/ theorem isInternal_submodule_iff_isCompl (A : ι → Submodule R M) {i j : ι} (hij : i ≠ j) (h : (Set.univ : Set ι) = {i, j}) : IsInternal A ↔ IsCompl (A i) (A j) := by have : ∀ k, k = i ∨ k = j := fun k ↦ by simpa using Set.ext_iff.mp h k - rw [isInternal_submodule_iff_independent_and_iSup_eq_top, iSup, ← Set.image_univ, h, - Set.image_insert_eq, Set.image_singleton, sSup_pair, CompleteLattice.independent_pair hij this] + rw [isInternal_submodule_iff_iSupIndep_and_iSup_eq_top, iSup, ← Set.image_univ, h, + Set.image_insert_eq, Set.image_singleton, sSup_pair, iSupIndep_pair hij this] exact ⟨fun ⟨hd, ht⟩ ↦ ⟨hd, codisjoint_iff.mpr ht⟩, fun ⟨hd, ht⟩ ↦ ⟨hd, ht.eq_top⟩⟩ @[simp] theorem isInternal_ne_bot_iff {A : ι → Submodule R M} : IsInternal (fun i : {i // A i ≠ ⊥} ↦ A i) ↔ IsInternal A := by - simp only [isInternal_submodule_iff_independent_and_iSup_eq_top] - exact Iff.and CompleteLattice.independent_ne_bot_iff_independent <| by simp + simp [isInternal_submodule_iff_iSupIndep_and_iSup_eq_top] -lemma isInternal_biSup_submodule_of_independent {A : ι → Submodule R M} (s : Set ι) - (h : CompleteLattice.Independent <| fun i : s ↦ A i) : +lemma isInternal_biSup_submodule_of_iSupIndep {A : ι → Submodule R M} (s : Set ι) + (h : iSupIndep <| fun i : s ↦ A i) : IsInternal <| fun (i : s) ↦ (A i).comap (⨆ i ∈ s, A i).subtype := by - refine (isInternal_submodule_iff_independent_and_iSup_eq_top _).mpr ⟨?_, by simp [iSup_subtype]⟩ + refine (isInternal_submodule_iff_iSupIndep_and_iSup_eq_top _).mpr ⟨?_, by simp [iSup_subtype]⟩ let p := ⨆ i ∈ s, A i have hp : ∀ i ∈ s, A i ≤ p := fun i hi ↦ le_biSup A hi let e : Submodule R p ≃o Set.Iic p := p.mapIic suffices (e ∘ fun i : s ↦ (A i).comap p.subtype) = fun i ↦ ⟨A i, hp i i.property⟩ by - rw [← CompleteLattice.independent_map_orderIso_iff e, this] - exact CompleteLattice.independent_of_independent_coe_Iic_comp h + rw [← iSupIndep_map_orderIso_iff e, this] + exact .of_coe_Iic_comp h ext i m change m ∈ ((A i).comap p.subtype).map p.subtype ↔ _ rw [Submodule.map_comap_subtype, inf_of_le_right (hp i i.property)] +@[deprecated (since := "2024-11-24")] +alias isInternal_biSup_submodule_of_independent := isInternal_biSup_submodule_of_iSupIndep + /-! Now copy the lemmas for subgroup and submonoids. -/ -theorem IsInternal.addSubmonoid_independent {M : Type*} [AddCommMonoid M] {A : ι → AddSubmonoid M} - (h : IsInternal A) : CompleteLattice.Independent A := - CompleteLattice.independent_of_dfinsupp_sumAddHom_injective _ h.injective +theorem IsInternal.addSubmonoid_iSupIndep {M : Type*} [AddCommMonoid M] {A : ι → AddSubmonoid M} + (h : IsInternal A) : iSupIndep A := + iSupIndep_of_dfinsupp_sumAddHom_injective _ h.injective + +@[deprecated (since := "2024-11-24")] +alias IsInternal.addSubmonoid_independent := IsInternal.addSubmonoid_iSupIndep + +theorem IsInternal.addSubgroup_iSupIndep {G : Type*} [AddCommGroup G] {A : ι → AddSubgroup G} + (h : IsInternal A) : iSupIndep A := + iSupIndep_of_dfinsupp_sumAddHom_injective' _ h.injective -theorem IsInternal.addSubgroup_independent {M : Type*} [AddCommGroup M] {A : ι → AddSubgroup M} - (h : IsInternal A) : CompleteLattice.Independent A := - CompleteLattice.independent_of_dfinsupp_sumAddHom_injective' _ h.injective +@[deprecated (since := "2024-11-24")] +alias IsInternal.addSubgroup_independent := IsInternal.addSubgroup_iSupIndep end Ring diff --git a/Mathlib/Algebra/DirectSum/Ring.lean b/Mathlib/Algebra/DirectSum/Ring.lean index aa5a84bec944c..883dc0bb086a7 100644 --- a/Mathlib/Algebra/DirectSum/Ring.lean +++ b/Mathlib/Algebra/DirectSum/Ring.lean @@ -63,9 +63,8 @@ instances for: * `A : ι → Submodule S`: `DirectSum.GSemiring.ofSubmodules`, `DirectSum.GCommSemiring.ofSubmodules`. -If `CompleteLattice.independent (Set.range A)`, these provide a gradation of `⨆ i, A i`, and the -mapping `⨁ i, A i →+ ⨆ i, A i` can be obtained as -`DirectSum.toMonoid (fun i ↦ AddSubmonoid.inclusion <| le_iSup A i)`. +If `sSupIndep A`, these provide a gradation of `⨆ i, A i`, and the mapping `⨁ i, A i →+ ⨆ i, A i` +can be obtained as `DirectSum.toMonoid (fun i ↦ AddSubmonoid.inclusion <| le_iSup A i)`. ## Tags diff --git a/Mathlib/Algebra/Lie/Semisimple/Basic.lean b/Mathlib/Algebra/Lie/Semisimple/Basic.lean index 12b30d31de1c2..50b8820a1addc 100644 --- a/Mathlib/Algebra/Lie/Semisimple/Basic.lean +++ b/Mathlib/Algebra/Lie/Semisimple/Basic.lean @@ -147,7 +147,7 @@ lemma isSimple_of_isAtom (I : LieIdeal R L) (hI : IsAtom I) : IsSimple R I where -- Finally `⁅b, y⁆ = 0`, by the independence of the atoms. · suffices ⁅b, y.val⁆ = 0 by erw [this]; simp only [zero_mem] rw [← LieSubmodule.mem_bot (R := R) (L := L), - ← (IsSemisimple.setIndependent_isAtom hI).eq_bot] + ← (IsSemisimple.sSupIndep_isAtom hI).eq_bot] exact ⟨lie_mem_right R L I b y y.2, lie_mem_left _ _ _ _ _ hb⟩ } -- Now that we know that `J` is an ideal of `L`, -- we start with the proof that `I` is a simple Lie algebra. @@ -240,7 +240,7 @@ lemma finitelyAtomistic : ∀ s : Finset (LieIdeal R L), ↑s ⊆ {I : LieIdeal constructor -- `j` brackets to `0` with `z`, since `⁅j, z⁆` is contained in `⁅J, K⁆ ≤ J ⊓ K`, -- and `J ⊓ K = ⊥` by the independence of the atoms. - · apply (setIndependent_isAtom.disjoint_sSup (hs hJs) hs'S (Finset.not_mem_erase _ _)).le_bot + · apply (sSupIndep_isAtom.disjoint_sSup (hs hJs) hs'S (Finset.not_mem_erase _ _)).le_bot apply LieSubmodule.lie_le_inf apply LieSubmodule.lie_mem_lie j.2 simpa only [K, Finset.sup_id_eq_sSup] using hz @@ -292,7 +292,7 @@ instance (priority := 100) IsSimple.instIsSemisimple [IsSimple R L] : IsSemisimple R L := by constructor · simp - · simpa using CompleteLattice.setIndependent_singleton _ + · simpa using sSupIndep_singleton _ · intro I hI₁ hI₂ apply IsSimple.non_abelian (R := R) (L := L) rw [IsSimple.isAtom_iff_eq_top] at hI₁ diff --git a/Mathlib/Algebra/Lie/Semisimple/Defs.lean b/Mathlib/Algebra/Lie/Semisimple/Defs.lean index 018f07cacb964..74db546a94643 100644 --- a/Mathlib/Algebra/Lie/Semisimple/Defs.lean +++ b/Mathlib/Algebra/Lie/Semisimple/Defs.lean @@ -72,7 +72,7 @@ class IsSemisimple : Prop where /-- In a semisimple Lie algebra, the supremum of the atoms is the whole Lie algebra. -/ sSup_atoms_eq_top : sSup {I : LieIdeal R L | IsAtom I} = ⊤ /-- In a semisimple Lie algebra, the atoms are independent. -/ - setIndependent_isAtom : CompleteLattice.SetIndependent {I : LieIdeal R L | IsAtom I} + sSupIndep_isAtom : sSupIndep {I : LieIdeal R L | IsAtom I} /-- In a semisimple Lie algebra, the atoms are non-abelian. -/ non_abelian_of_isAtom : ∀ I : LieIdeal R L, IsAtom I → ¬ IsLieAbelian I diff --git a/Mathlib/Algebra/Lie/Submodule.lean b/Mathlib/Algebra/Lie/Submodule.lean index c6fc1d659955f..1201232787862 100644 --- a/Mathlib/Algebra/Lie/Submodule.lean +++ b/Mathlib/Algebra/Lie/Submodule.lean @@ -514,9 +514,12 @@ theorem isCompl_iff_coe_toSubmodule : IsCompl N N' ↔ IsCompl (N : Submodule R M) (N' : Submodule R M) := by simp only [isCompl_iff, disjoint_iff_coe_toSubmodule, codisjoint_iff_coe_toSubmodule] -theorem independent_iff_coe_toSubmodule {ι : Type*} {N : ι → LieSubmodule R L M} : - CompleteLattice.Independent N ↔ CompleteLattice.Independent fun i ↦ (N i : Submodule R M) := by - simp [CompleteLattice.independent_def, disjoint_iff_coe_toSubmodule] +theorem iSupIndep_iff_coe_toSubmodule {ι : Type*} {N : ι → LieSubmodule R L M} : + iSupIndep N ↔ iSupIndep fun i ↦ (N i : Submodule R M) := by + simp [iSupIndep_def, disjoint_iff_coe_toSubmodule] + +@[deprecated (since := "2024-11-24")] +alias independent_iff_coe_toSubmodule := iSupIndep_iff_coe_toSubmodule theorem iSup_eq_top_iff_coe_toSubmodule {ι : Sort*} {N : ι → LieSubmodule R L M} : ⨆ i, N i = ⊤ ↔ ⨆ i, (N i : Submodule R M) = ⊤ := by diff --git a/Mathlib/Algebra/Lie/TraceForm.lean b/Mathlib/Algebra/Lie/TraceForm.lean index 2d1253e2b18c9..55ee3be037f29 100644 --- a/Mathlib/Algebra/Lie/TraceForm.lean +++ b/Mathlib/Algebra/Lie/TraceForm.lean @@ -227,8 +227,8 @@ lemma traceForm_eq_sum_genWeightSpaceOf convert finite_genWeightSpaceOf_ne_bot R L M z exact LieSubmodule.coeSubmodule_eq_bot_iff (genWeightSpaceOf M _ _) classical - have h := LieSubmodule.independent_iff_coe_toSubmodule.mp <| independent_genWeightSpaceOf R L M z - have hds := DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top h <| by + have h := LieSubmodule.iSupIndep_iff_coe_toSubmodule.mp <| iSupIndep_genWeightSpaceOf R L M z + have hds := DirectSum.isInternal_submodule_of_iSupIndep_of_iSup_eq_top h <| by simp [← LieSubmodule.iSup_coe_toSubmodule] simp only [LinearMap.coeFn_sum, Finset.sum_apply, traceForm_apply_apply, LinearMap.trace_eq_sum_trace_restrict' hds hfin hxy] @@ -407,8 +407,8 @@ lemma traceForm_eq_sum_finrank_nsmul_mul (x y : L) : (genWeightSpace M χ) (genWeightSpace M χ) := fun χ m hm ↦ LieSubmodule.lie_mem _ <| LieSubmodule.lie_mem _ hm classical - have hds := DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top - (LieSubmodule.independent_iff_coe_toSubmodule.mp <| independent_genWeightSpace' K L M) + have hds := DirectSum.isInternal_submodule_of_iSupIndep_of_iSup_eq_top + (LieSubmodule.iSupIndep_iff_coe_toSubmodule.mp <| iSupIndep_genWeightSpace' K L M) (LieSubmodule.iSup_eq_top_iff_coe_toSubmodule.mp <| iSup_genWeightSpace_eq_top' K L M) simp_rw [traceForm_apply_apply, LinearMap.trace_eq_sum_trace_restrict hds hxy, ← traceForm_genWeightSpace_eq K L M _ x y] diff --git a/Mathlib/Algebra/Lie/Weights/Basic.lean b/Mathlib/Algebra/Lie/Weights/Basic.lean index ce2e34ec7f66e..8ddf44e6372c7 100644 --- a/Mathlib/Algebra/Lie/Weights/Basic.lean +++ b/Mathlib/Algebra/Lie/Weights/Basic.lean @@ -33,7 +33,7 @@ Basic definitions and properties of the above ideas are provided in this file. * `LieModule.iSup_ucs_eq_genWeightSpace_zero` * `LieModule.iInf_lowerCentralSeries_eq_posFittingComp` * `LieModule.isCompl_genWeightSpace_zero_posFittingComp` - * `LieModule.independent_genWeightSpace` + * `LieModule.iSupIndep_genWeightSpace` * `LieModule.iSup_genWeightSpace_eq_top` ## References @@ -666,32 +666,39 @@ lemma injOn_genWeightSpace [NoZeroSMulDivisors R M] : /-- Lie module weight spaces are independent. -See also `LieModule.independent_genWeightSpace'`. -/ -lemma independent_genWeightSpace [NoZeroSMulDivisors R M] : - CompleteLattice.Independent fun χ : L → R ↦ genWeightSpace M χ := by - simp only [LieSubmodule.independent_iff_coe_toSubmodule, genWeightSpace, +See also `LieModule.iSupIndep_genWeightSpace'`. -/ +lemma iSupIndep_genWeightSpace [NoZeroSMulDivisors R M] : + iSupIndep fun χ : L → R ↦ genWeightSpace M χ := by + simp only [LieSubmodule.iSupIndep_iff_coe_toSubmodule, genWeightSpace, LieSubmodule.iInf_coe_toSubmodule] exact Module.End.independent_iInf_maxGenEigenspace_of_forall_mapsTo (toEnd R L M) (fun x y φ z ↦ (genWeightSpaceOf M φ y).lie_mem) -lemma independent_genWeightSpace' [NoZeroSMulDivisors R M] : - CompleteLattice.Independent fun χ : Weight R L M ↦ genWeightSpace M χ := - (independent_genWeightSpace R L M).comp <| +@[deprecated (since := "2024-11-24")] alias independent_genWeightSpace := iSupIndep_genWeightSpace + +lemma iSupIndep_genWeightSpace' [NoZeroSMulDivisors R M] : + iSupIndep fun χ : Weight R L M ↦ genWeightSpace M χ := + (iSupIndep_genWeightSpace R L M).comp <| Subtype.val_injective.comp (Weight.equivSetOf R L M).injective -lemma independent_genWeightSpaceOf [NoZeroSMulDivisors R M] (x : L) : - CompleteLattice.Independent fun (χ : R) ↦ genWeightSpaceOf M χ x := by - rw [LieSubmodule.independent_iff_coe_toSubmodule] +@[deprecated (since := "2024-11-24")] alias independent_genWeightSpace' := iSupIndep_genWeightSpace' + +lemma iSupIndep_genWeightSpaceOf [NoZeroSMulDivisors R M] (x : L) : + iSupIndep fun (χ : R) ↦ genWeightSpaceOf M χ x := by + rw [LieSubmodule.iSupIndep_iff_coe_toSubmodule] dsimp [genWeightSpaceOf] exact (toEnd R L M x).independent_genEigenspace _ +@[deprecated (since := "2024-11-24")] +alias independent_genWeightSpaceOf := iSupIndep_genWeightSpaceOf + lemma finite_genWeightSpaceOf_ne_bot [NoZeroSMulDivisors R M] [IsNoetherian R M] (x : L) : {χ : R | genWeightSpaceOf M χ x ≠ ⊥}.Finite := - CompleteLattice.WellFoundedGT.finite_ne_bot_of_independent (independent_genWeightSpaceOf R L M x) + WellFoundedGT.finite_ne_bot_of_iSupIndep (iSupIndep_genWeightSpaceOf R L M x) lemma finite_genWeightSpace_ne_bot [NoZeroSMulDivisors R M] [IsNoetherian R M] : {χ : L → R | genWeightSpace M χ ≠ ⊥}.Finite := - CompleteLattice.WellFoundedGT.finite_ne_bot_of_independent (independent_genWeightSpace R L M) + WellFoundedGT.finite_ne_bot_of_iSupIndep (iSupIndep_genWeightSpace R L M) instance Weight.instFinite [NoZeroSMulDivisors R M] [IsNoetherian R M] : Finite (Weight R L M) := by diff --git a/Mathlib/Algebra/Lie/Weights/Chain.lean b/Mathlib/Algebra/Lie/Weights/Chain.lean index 4107e1e4cd27f..1acfd102224c7 100644 --- a/Mathlib/Algebra/Lie/Weights/Chain.lean +++ b/Mathlib/Algebra/Lie/Weights/Chain.lean @@ -204,9 +204,9 @@ lemma exists_forall_mem_corootSpace_smul_add_eq_zero exact finrank_pos refine ⟨a, b, Int.ofNat_pos.mpr hb, fun x hx ↦ ?_⟩ let N : ℤ → Submodule R M := fun k ↦ genWeightSpace M (k • α + χ) - have h₁ : CompleteLattice.Independent fun (i : Finset.Ioo p q) ↦ N i := by - rw [← LieSubmodule.independent_iff_coe_toSubmodule] - refine (independent_genWeightSpace R H M).comp fun i j hij ↦ ?_ + have h₁ : iSupIndep fun (i : Finset.Ioo p q) ↦ N i := by + rw [← LieSubmodule.iSupIndep_iff_coe_toSubmodule] + refine (iSupIndep_genWeightSpace R H M).comp fun i j hij ↦ ?_ exact SetCoe.ext <| smul_left_injective ℤ hα <| by rwa [add_left_inj] at hij have h₂ : ∀ i, MapsTo (toEnd R H M x) ↑(N i) ↑(N i) := fun _ _ ↦ LieSubmodule.lie_mem _ have h₃ : genWeightSpaceChain M α χ p q = ⨆ i ∈ Finset.Ioo p q, N i := by diff --git a/Mathlib/Algebra/Lie/Weights/Killing.lean b/Mathlib/Algebra/Lie/Weights/Killing.lean index 3fe00e4149748..7c2b6ece5a6b5 100644 --- a/Mathlib/Algebra/Lie/Weights/Killing.lean +++ b/Mathlib/Algebra/Lie/Weights/Killing.lean @@ -112,8 +112,8 @@ lemma killingForm_apply_eq_zero_of_mem_rootSpace_of_add_ne_zero {α β : H → K (mapsTo_toEnd_genWeightSpace_add_of_mem_rootSpace K L H L α (β + γ) hx).comp <| mapsTo_toEnd_genWeightSpace_add_of_mem_rootSpace K L H L β γ hy classical - have hds := DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top - (LieSubmodule.independent_iff_coe_toSubmodule.mp <| independent_genWeightSpace K H L) + have hds := DirectSum.isInternal_submodule_of_iSupIndep_of_iSup_eq_top + (LieSubmodule.iSupIndep_iff_coe_toSubmodule.mp <| iSupIndep_genWeightSpace K H L) (LieSubmodule.iSup_eq_top_iff_coe_toSubmodule.mp <| iSup_genWeightSpace_eq_top K H L) exact LinearMap.trace_eq_zero_of_mapsTo_ne hds σ hσ hf diff --git a/Mathlib/Algebra/Module/Torsion.lean b/Mathlib/Algebra/Module/Torsion.lean index 3121a844257b9..08b33bcef771a 100644 --- a/Mathlib/Algebra/Module/Torsion.lean +++ b/Mathlib/Algebra/Module/Torsion.lean @@ -96,13 +96,13 @@ theorem torsionOf_eq_bot_iff_of_noZeroSMulDivisors [Nontrivial R] [NoZeroSMulDiv · rw [mem_torsionOf_iff, smul_eq_zero] at hr tauto -/-- See also `CompleteLattice.Independent.linearIndependent` which provides the same conclusion +/-- See also `iSupIndep.linearIndependent` which provides the same conclusion but requires the stronger hypothesis `NoZeroSMulDivisors R M`. -/ -theorem CompleteLattice.Independent.linear_independent' {ι R M : Type*} {v : ι → M} [Ring R] - [AddCommGroup M] [Module R M] (hv : CompleteLattice.Independent fun i => R ∙ v i) +theorem iSupIndep.linearIndependent' {ι R M : Type*} {v : ι → M} [Ring R] + [AddCommGroup M] [Module R M] (hv : iSupIndep fun i => R ∙ v i) (h_ne_zero : ∀ i, Ideal.torsionOf R M (v i) = ⊥) : LinearIndependent R v := by refine linearIndependent_iff_not_smul_mem_span.mpr fun i r hi => ?_ - replace hv := CompleteLattice.independent_def.mp hv i + replace hv := iSupIndep_def.mp hv i simp only [iSup_subtype', ← Submodule.span_range_eq_iSup (ι := Subtype _), disjoint_iff] at hv have : r • v i ∈ (⊥ : Submodule R M) := by rw [← hv, Submodule.mem_inf] @@ -113,6 +113,9 @@ theorem CompleteLattice.Independent.linear_independent' {ι R M : Type*} {v : ι rw [← Submodule.mem_bot R, ← h_ne_zero i] simpa using this +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.Independent.linear_independent' := iSupIndep.linearIndependent' + end TorsionOf section @@ -443,8 +446,8 @@ theorem torsionBySet_isInternal {p : ι → Ideal R} (hp : (S : Set ι).Pairwise fun i j => p i ⊔ p j = ⊤) (hM : Module.IsTorsionBySet R M (⨅ i ∈ S, p i : Ideal R)) : DirectSum.IsInternal fun i : S => torsionBySet R M <| p i := - DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top - (CompleteLattice.independent_iff_supIndep.mpr <| supIndep_torsionBySet_ideal hp) + DirectSum.isInternal_submodule_of_iSupIndep_of_iSup_eq_top + (iSupIndep_iff_supIndep.mpr <| supIndep_torsionBySet_ideal hp) (by apply (iSup_subtype'' ↑S fun i => torsionBySet R M <| p i).trans -- Porting note: times out if we change apply below to <| diff --git a/Mathlib/Analysis/InnerProductSpace/Basic.lean b/Mathlib/Analysis/InnerProductSpace/Basic.lean index 4acfe80398b10..2f715eb17ed19 100644 --- a/Mathlib/Analysis/InnerProductSpace/Basic.lean +++ b/Mathlib/Analysis/InnerProductSpace/Basic.lean @@ -2197,9 +2197,9 @@ elements each from a different subspace in the family is linearly independent. I pairwise intersections of elements of the family are 0. -/ theorem OrthogonalFamily.independent {V : ι → Submodule 𝕜 E} (hV : OrthogonalFamily 𝕜 (fun i => V i) fun i => (V i).subtypeₗᵢ) : - CompleteLattice.Independent V := by + iSupIndep V := by classical - apply CompleteLattice.independent_of_dfinsupp_lsum_injective + apply iSupIndep_of_dfinsupp_lsum_injective refine LinearMap.ker_eq_bot.mp ?_ rw [Submodule.eq_bot_iff] intro v hv diff --git a/Mathlib/Analysis/InnerProductSpace/Projection.lean b/Mathlib/Analysis/InnerProductSpace/Projection.lean index c2ea3c5347107..6be13bb9ce5f1 100644 --- a/Mathlib/Analysis/InnerProductSpace/Projection.lean +++ b/Mathlib/Analysis/InnerProductSpace/Projection.lean @@ -1218,7 +1218,7 @@ theorem OrthogonalFamily.isInternal_iff_of_isComplete [DecidableEq ι] {V : ι (hV : OrthogonalFamily 𝕜 (fun i => V i) fun i => (V i).subtypeₗᵢ) (hc : IsComplete (↑(iSup V) : Set E)) : DirectSum.IsInternal V ↔ (iSup V)ᗮ = ⊥ := by haveI : CompleteSpace (↥(iSup V)) := hc.completeSpace_coe - simp only [DirectSum.isInternal_submodule_iff_independent_and_iSup_eq_top, hV.independent, + simp only [DirectSum.isInternal_submodule_iff_iSupIndep_and_iSup_eq_top, hV.independent, true_and, Submodule.orthogonal_eq_bot_iff] /-- An orthogonal family of subspaces of `E` satisfies `DirectSum.IsInternal` (that is, diff --git a/Mathlib/GroupTheory/NoncommPiCoprod.lean b/Mathlib/GroupTheory/NoncommPiCoprod.lean index 3e22ae8416ac7..e5b14b7d09118 100644 --- a/Mathlib/GroupTheory/NoncommPiCoprod.lean +++ b/Mathlib/GroupTheory/NoncommPiCoprod.lean @@ -30,7 +30,7 @@ images of different morphisms commute, we obtain a canonical morphism * `MonoidHom.noncommPiCoprod_range`: The range of `MonoidHom.noncommPiCoprod` is `⨆ (i : ι), (ϕ i).range` * `Subgroup.noncommPiCoprod_range`: The range of `Subgroup.noncommPiCoprod` is `⨆ (i : ι), H i`. -* `MonoidHom.injective_noncommPiCoprod_of_independent`: in the case of groups, `pi_hom.hom` is +* `MonoidHom.injective_noncommPiCoprod_of_iSupIndep`: in the case of groups, `pi_hom.hom` is injective if the `ϕ` are injective and the ranges of the `ϕ` are independent. * `MonoidHom.independent_range_of_coprime_order`: If the `N i` have coprime orders, then the ranges of the `ϕ` are independent. @@ -48,8 +48,8 @@ variable {G : Type*} [Group G] generalizes (one direction of) `Subgroup.disjoint_iff_mul_eq_one`. -/ @[to_additive "`Finset.noncommSum` is “injective” in `f` if `f` maps into independent subgroups. This generalizes (one direction of) `AddSubgroup.disjoint_iff_add_eq_zero`. "] -theorem eq_one_of_noncommProd_eq_one_of_independent {ι : Type*} (s : Finset ι) (f : ι → G) (comm) - (K : ι → Subgroup G) (hind : CompleteLattice.Independent K) (hmem : ∀ x ∈ s, f x ∈ K x) +theorem eq_one_of_noncommProd_eq_one_of_iSupIndep {ι : Type*} (s : Finset ι) (f : ι → G) (comm) + (K : ι → Subgroup G) (hind : iSupIndep K) (hmem : ∀ x ∈ s, f x ∈ K x) (heq1 : s.noncommProd f comm = 1) : ∀ i ∈ s, f i = 1 := by classical revert heq1 @@ -73,6 +73,9 @@ theorem eq_one_of_noncommProd_eq_one_of_independent {ι : Type*} (s : Finset ι) · exact heq1i · refine ih hcomm hmem.2 heq1S _ h +@[deprecated (since := "2024-11-24")] +alias eq_one_of_noncommProd_eq_one_of_independent := eq_one_of_noncommProd_eq_one_of_iSupIndep + end Subgroup section FamilyOfMonoids @@ -200,27 +203,30 @@ theorem noncommPiCoprod_range [Fintype ι] exact ⟨Pi.mulSingle i y, noncommPiCoprod_mulSingle _ _ _⟩ @[to_additive] -theorem injective_noncommPiCoprod_of_independent [Fintype ι] +theorem injective_noncommPiCoprod_of_iSupIndep [Fintype ι] {hcomm : Pairwise fun i j : ι => ∀ (x : H i) (y : H j), Commute (ϕ i x) (ϕ j y)} - (hind : CompleteLattice.Independent fun i => (ϕ i).range) + (hind : iSupIndep fun i => (ϕ i).range) (hinj : ∀ i, Function.Injective (ϕ i)) : Function.Injective (noncommPiCoprod ϕ hcomm) := by classical apply (MonoidHom.ker_eq_bot_iff _).mp rw [eq_bot_iff] intro f heq1 have : ∀ i, i ∈ Finset.univ → ϕ i (f i) = 1 := - Subgroup.eq_one_of_noncommProd_eq_one_of_independent _ _ (fun _ _ _ _ h => hcomm h _ _) + Subgroup.eq_one_of_noncommProd_eq_one_of_iSupIndep _ _ (fun _ _ _ _ h => hcomm h _ _) _ hind (by simp) heq1 ext i apply hinj simp [this i (Finset.mem_univ i)] +@[deprecated (since := "2024-11-24")] +alias injective_noncommPiCoprod_of_independent := injective_noncommPiCoprod_of_iSupIndep + @[to_additive] theorem independent_range_of_coprime_order (hcomm : Pairwise fun i j : ι => ∀ (x : H i) (y : H j), Commute (ϕ i x) (ϕ j y)) [Finite ι] [∀ i, Fintype (H i)] (hcoprime : Pairwise fun i j => Nat.Coprime (Fintype.card (H i)) (Fintype.card (H j))) : - CompleteLattice.Independent fun i => (ϕ i).range := by + iSupIndep fun i => (ϕ i).range := by cases nonempty_fintype ι letI := Classical.decEq ι rintro i @@ -279,7 +285,7 @@ theorem independent_of_coprime_order (hcomm : Pairwise fun i j : ι => ∀ x y : G, x ∈ H i → y ∈ H j → Commute x y) [Finite ι] [∀ i, Fintype (H i)] (hcoprime : Pairwise fun i j => Nat.Coprime (Fintype.card (H i)) (Fintype.card (H j))) : - CompleteLattice.Independent H := by + iSupIndep H := by simpa using MonoidHom.independent_range_of_coprime_order (fun i => (H i).subtype) (commute_subtype_of_commute hcomm) hcoprime @@ -306,11 +312,11 @@ theorem noncommPiCoprod_range simp [noncommPiCoprod, MonoidHom.noncommPiCoprod_range] @[to_additive] -theorem injective_noncommPiCoprod_of_independent +theorem injective_noncommPiCoprod_of_iSupIndep {hcomm : Pairwise fun i j : ι => ∀ x y : G, x ∈ H i → y ∈ H j → Commute x y} - (hind : CompleteLattice.Independent H) : + (hind : iSupIndep H) : Function.Injective (noncommPiCoprod hcomm) := by - apply MonoidHom.injective_noncommPiCoprod_of_independent + apply MonoidHom.injective_noncommPiCoprod_of_iSupIndep · simpa using hind · intro i exact Subtype.coe_injective diff --git a/Mathlib/GroupTheory/Sylow.lean b/Mathlib/GroupTheory/Sylow.lean index 5521719559622..fdfa524a631b9 100644 --- a/Mathlib/GroupTheory/Sylow.lean +++ b/Mathlib/GroupTheory/Sylow.lean @@ -806,7 +806,7 @@ noncomputable def directProductOfNormal [Finite G] apply MulEquiv.ofBijective (Subgroup.noncommPiCoprod hcomm) apply (Fintype.bijective_iff_injective_and_card _).mpr constructor - · apply Subgroup.injective_noncommPiCoprod_of_independent + · apply Subgroup.injective_noncommPiCoprod_of_iSupIndep apply independent_of_coprime_order hcomm rintro ⟨p₁, hp₁⟩ ⟨p₂, hp₂⟩ hne haveI hp₁' := Fact.mk (Nat.prime_of_mem_primeFactors hp₁) diff --git a/Mathlib/LinearAlgebra/DFinsupp.lean b/Mathlib/LinearAlgebra/DFinsupp.lean index 406aa08a4c032..f411c6bc0c1b5 100644 --- a/Mathlib/LinearAlgebra/DFinsupp.lean +++ b/Mathlib/LinearAlgebra/DFinsupp.lean @@ -390,8 +390,6 @@ theorem mem_iSup_finset_iff_exists_sum {s : Finset ι} (p : ι → Submodule R N end Submodule -namespace CompleteLattice - open DFinsupp section Semiring @@ -401,23 +399,26 @@ variable [DecidableEq ι] [Semiring R] [AddCommMonoid N] [Module R N] /-- Independence of a family of submodules can be expressed as a quantifier over `DFinsupp`s. This is an intermediate result used to prove -`CompleteLattice.independent_of_dfinsupp_lsum_injective` and -`CompleteLattice.Independent.dfinsupp_lsum_injective`. -/ -theorem independent_iff_forall_dfinsupp (p : ι → Submodule R N) : - Independent p ↔ +`iSupIndep_of_dfinsupp_lsum_injective` and +`iSupIndep.dfinsupp_lsum_injective`. -/ +theorem iSupIndep_iff_forall_dfinsupp (p : ι → Submodule R N) : + iSupIndep p ↔ ∀ (i) (x : p i) (v : Π₀ i : ι, ↥(p i)), lsum ℕ (M := fun i ↦ ↥(p i)) (fun i => (p i).subtype) (erase i v) = x → x = 0 := by - simp_rw [CompleteLattice.independent_def, Submodule.disjoint_def, + simp_rw [iSupIndep_def, Submodule.disjoint_def, Submodule.mem_biSup_iff_exists_dfinsupp, exists_imp, filter_ne_eq_erase] refine forall_congr' fun i => Subtype.forall'.trans ?_ simp_rw [Submodule.coe_eq_zero] +@[deprecated (since := "2024-11-24")] +alias independent_iff_forall_dfinsupp := iSupIndep_iff_forall_dfinsupp + /- If `DFinsupp.lsum` applied with `Submodule.subtype` is injective then the submodules are -independent. -/ -theorem independent_of_dfinsupp_lsum_injective (p : ι → Submodule R N) +iSupIndep. -/ +theorem iSupIndep_of_dfinsupp_lsum_injective (p : ι → Submodule R N) (h : Function.Injective (lsum ℕ (M := fun i ↦ ↥(p i)) fun i => (p i).subtype)) : - Independent p := by - rw [independent_iff_forall_dfinsupp] + iSupIndep p := by + rw [iSupIndep_iff_forall_dfinsupp] intro i x v hv replace hv : lsum ℕ (M := fun i ↦ ↥(p i)) (fun i => (p i).subtype) (erase i v) = lsum ℕ (M := fun i ↦ ↥(p i)) (fun i => (p i).subtype) (single i x) := by @@ -425,12 +426,18 @@ theorem independent_of_dfinsupp_lsum_injective (p : ι → Submodule R N) have := DFunLike.ext_iff.mp (h hv) i simpa [eq_comm] using this +@[deprecated (since := "2024-11-24")] +alias independent_of_dfinsupp_lsum_injective := iSupIndep_of_dfinsupp_lsum_injective + /- If `DFinsupp.sumAddHom` applied with `AddSubmonoid.subtype` is injective then the additive submonoids are independent. -/ -theorem independent_of_dfinsupp_sumAddHom_injective (p : ι → AddSubmonoid N) - (h : Function.Injective (sumAddHom fun i => (p i).subtype)) : Independent p := by - rw [← independent_map_orderIso_iff (AddSubmonoid.toNatSubmodule : AddSubmonoid N ≃o _)] - exact independent_of_dfinsupp_lsum_injective _ h +theorem iSupIndep_of_dfinsupp_sumAddHom_injective (p : ι → AddSubmonoid N) + (h : Function.Injective (sumAddHom fun i => (p i).subtype)) : iSupIndep p := by + rw [← iSupIndep_map_orderIso_iff (AddSubmonoid.toNatSubmodule : AddSubmonoid N ≃o _)] + exact iSupIndep_of_dfinsupp_lsum_injective _ h + +@[deprecated (since := "2024-11-24")] +alias independent_of_dfinsupp_sumAddHom_injective := iSupIndep_of_dfinsupp_sumAddHom_injective /-- Combining `DFinsupp.lsum` with `LinearMap.toSpanSingleton` is the same as `Finsupp.linearCombination` -/ @@ -450,24 +457,27 @@ section Ring variable [DecidableEq ι] [Ring R] [AddCommGroup N] [Module R N] -/- If `DFinsupp.sumAddHom` applied with `AddSubmonoid.subtype` is injective then the additive +/-- If `DFinsupp.sumAddHom` applied with `AddSubmonoid.subtype` is injective then the additive subgroups are independent. -/ -theorem independent_of_dfinsupp_sumAddHom_injective' (p : ι → AddSubgroup N) - (h : Function.Injective (sumAddHom fun i => (p i).subtype)) : Independent p := by - rw [← independent_map_orderIso_iff (AddSubgroup.toIntSubmodule : AddSubgroup N ≃o _)] - exact independent_of_dfinsupp_lsum_injective _ h +theorem iSupIndep_of_dfinsupp_sumAddHom_injective' (p : ι → AddSubgroup N) + (h : Function.Injective (sumAddHom fun i => (p i).subtype)) : iSupIndep p := by + rw [← iSupIndep_map_orderIso_iff (AddSubgroup.toIntSubmodule : AddSubgroup N ≃o _)] + exact iSupIndep_of_dfinsupp_lsum_injective _ h + +@[deprecated (since := "2024-11-24")] +alias independent_of_dfinsupp_sumAddHom_injective' := iSupIndep_of_dfinsupp_sumAddHom_injective' /-- The canonical map out of a direct sum of a family of submodules is injective when the submodules -are `CompleteLattice.Independent`. +are `iSupIndep`. Note that this is not generally true for `[Semiring R]`, for instance when `A` is the `ℕ`-submodules of the positive and negative integers. See `Counterexamples/DirectSumIsInternal.lean` for a proof of this fact. -/ -theorem Independent.dfinsupp_lsum_injective {p : ι → Submodule R N} (h : Independent p) : +theorem iSupIndep.dfinsupp_lsum_injective {p : ι → Submodule R N} (h : iSupIndep p) : Function.Injective (lsum ℕ (M := fun i ↦ ↥(p i)) fun i => (p i).subtype) := by -- simplify everything down to binders over equalities in `N` - rw [independent_iff_forall_dfinsupp] at h + rw [iSupIndep_iff_forall_dfinsupp] at h suffices LinearMap.ker (lsum ℕ (M := fun i ↦ ↥(p i)) fun i => (p i).subtype) = ⊥ by -- Lean can't find this without our help letI thisI : AddCommGroup (Π₀ i, p i) := inferInstance @@ -482,34 +492,46 @@ theorem Independent.dfinsupp_lsum_injective {p : ι → Submodule R N} (h : Inde rwa [← erase_add_single i m, LinearMap.map_add, lsum_single, Submodule.subtype_apply, add_eq_zero_iff_eq_neg, ← Submodule.coe_neg] at hm +@[deprecated (since := "2024-11-24")] +alias Independent.dfinsupp_lsum_injective := iSupIndep.dfinsupp_lsum_injective + /-- The canonical map out of a direct sum of a family of additive subgroups is injective when the -additive subgroups are `CompleteLattice.Independent`. -/ -theorem Independent.dfinsupp_sumAddHom_injective {p : ι → AddSubgroup N} (h : Independent p) : +additive subgroups are `iSupIndep`. -/ +theorem iSupIndep.dfinsupp_sumAddHom_injective {p : ι → AddSubgroup N} (h : iSupIndep p) : Function.Injective (sumAddHom fun i => (p i).subtype) := by - rw [← independent_map_orderIso_iff (AddSubgroup.toIntSubmodule : AddSubgroup N ≃o _)] at h + rw [← iSupIndep_map_orderIso_iff (AddSubgroup.toIntSubmodule : AddSubgroup N ≃o _)] at h exact h.dfinsupp_lsum_injective +@[deprecated (since := "2024-11-24")] +alias Independent.dfinsupp_sumAddHom_injective := iSupIndep.dfinsupp_sumAddHom_injective + /-- A family of submodules over an additive group are independent if and only iff `DFinsupp.lsum` applied with `Submodule.subtype` is injective. Note that this is not generally true for `[Semiring R]`; see -`CompleteLattice.Independent.dfinsupp_lsum_injective` for details. -/ -theorem independent_iff_dfinsupp_lsum_injective (p : ι → Submodule R N) : - Independent p ↔ Function.Injective (lsum ℕ (M := fun i ↦ ↥(p i)) fun i => (p i).subtype) := - ⟨Independent.dfinsupp_lsum_injective, independent_of_dfinsupp_lsum_injective p⟩ +`iSupIndep.dfinsupp_lsum_injective` for details. -/ +theorem iSupIndep_iff_dfinsupp_lsum_injective (p : ι → Submodule R N) : + iSupIndep p ↔ Function.Injective (lsum ℕ (M := fun i ↦ ↥(p i)) fun i => (p i).subtype) := + ⟨iSupIndep.dfinsupp_lsum_injective, iSupIndep_of_dfinsupp_lsum_injective p⟩ + +@[deprecated (since := "2024-11-24")] +alias independent_iff_dfinsupp_lsum_injective := iSupIndep_iff_dfinsupp_lsum_injective /-- A family of additive subgroups over an additive group are independent if and only if `DFinsupp.sumAddHom` applied with `AddSubgroup.subtype` is injective. -/ -theorem independent_iff_dfinsupp_sumAddHom_injective (p : ι → AddSubgroup N) : - Independent p ↔ Function.Injective (sumAddHom fun i => (p i).subtype) := - ⟨Independent.dfinsupp_sumAddHom_injective, independent_of_dfinsupp_sumAddHom_injective' p⟩ +theorem iSupIndep_iff_dfinsupp_sumAddHom_injective (p : ι → AddSubgroup N) : + iSupIndep p ↔ Function.Injective (sumAddHom fun i => (p i).subtype) := + ⟨iSupIndep.dfinsupp_sumAddHom_injective, iSupIndep_of_dfinsupp_sumAddHom_injective' p⟩ + +@[deprecated (since := "2024-11-24")] +alias independent_iff_dfinsupp_sumAddHom_injective := iSupIndep_iff_dfinsupp_sumAddHom_injective -/-- If a family of submodules is `Independent`, then a choice of nonzero vector from each submodule +/-- If a family of submodules is independent, then a choice of nonzero vector from each submodule forms a linearly independent family. -See also `CompleteLattice.Independent.linearIndependent'`. -/ -theorem Independent.linearIndependent [NoZeroSMulDivisors R N] {ι} (p : ι → Submodule R N) - (hp : Independent p) {v : ι → N} (hv : ∀ i, v i ∈ p i) (hv' : ∀ i, v i ≠ 0) : +See also `iSupIndep.linearIndependent'`. -/ +theorem iSupIndep.linearIndependent [NoZeroSMulDivisors R N] {ι} (p : ι → Submodule R N) + (hp : iSupIndep p) {v : ι → N} (hv : ∀ i, v i ∈ p i) (hv' : ∀ i, v i ≠ 0) : LinearIndependent R v := by let _ := Classical.decEq ι let _ := Classical.decEq R @@ -527,15 +549,19 @@ theorem Independent.linearIndependent [NoZeroSMulDivisors R N] {ι} (p : ι → simp only [coe_zero, Pi.zero_apply, ZeroMemClass.coe_zero, smul_eq_zero, ha] at this simpa -theorem independent_iff_linearIndependent_of_ne_zero [NoZeroSMulDivisors R N] {ι} {v : ι → N} - (h_ne_zero : ∀ i, v i ≠ 0) : (Independent fun i => R ∙ v i) ↔ LinearIndependent R v := +@[deprecated (since := "2024-11-24")] +alias Independent.linearIndependent := iSupIndep.linearIndependent + +theorem iSupIndep_iff_linearIndependent_of_ne_zero [NoZeroSMulDivisors R N] {ι} {v : ι → N} + (h_ne_zero : ∀ i, v i ≠ 0) : (iSupIndep fun i => R ∙ v i) ↔ LinearIndependent R v := let _ := Classical.decEq ι ⟨fun hv => hv.linearIndependent _ (fun i => Submodule.mem_span_singleton_self <| v i) h_ne_zero, - fun hv => hv.independent_span_singleton⟩ + fun hv => hv.iSupIndep_span_singleton⟩ -end Ring +@[deprecated (since := "2024-11-24")] +alias independent_iff_linearIndependent_of_ne_zero := iSupIndep_iff_linearIndependent_of_ne_zero -end CompleteLattice +end Ring namespace LinearMap diff --git a/Mathlib/LinearAlgebra/Dimension/Finite.lean b/Mathlib/LinearAlgebra/Dimension/Finite.lean index 49fba6493d2d1..6c000a000aaf9 100644 --- a/Mathlib/LinearAlgebra/Dimension/Finite.lean +++ b/Mathlib/LinearAlgebra/Dimension/Finite.lean @@ -249,8 +249,8 @@ theorem Module.Finite.not_linearIndependent_of_infinite {ι : Type*} [Infinite section variable [NoZeroSMulDivisors R M] -theorem CompleteLattice.Independent.subtype_ne_bot_le_rank [Nontrivial R] - {V : ι → Submodule R M} (hV : CompleteLattice.Independent V) : +theorem iSupIndep.subtype_ne_bot_le_rank [Nontrivial R] + {V : ι → Submodule R M} (hV : iSupIndep V) : Cardinal.lift.{v} #{ i : ι // V i ≠ ⊥ } ≤ Cardinal.lift.{w} (Module.rank R M) := by set I := { i : ι // V i ≠ ⊥ } have hI : ∀ i : I, ∃ v ∈ V i, v ≠ (0 : M) := by @@ -261,10 +261,13 @@ theorem CompleteLattice.Independent.subtype_ne_bot_le_rank [Nontrivial R] have : LinearIndependent R v := (hV.comp Subtype.coe_injective).linearIndependent _ hvV hv exact this.cardinal_lift_le_rank +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.Independent.subtype_ne_bot_le_rank := iSupIndep.subtype_ne_bot_le_rank + variable [Module.Finite R M] [StrongRankCondition R] -theorem CompleteLattice.Independent.subtype_ne_bot_le_finrank_aux - {p : ι → Submodule R M} (hp : CompleteLattice.Independent p) : +theorem iSupIndep.subtype_ne_bot_le_finrank_aux + {p : ι → Submodule R M} (hp : iSupIndep p) : #{ i // p i ≠ ⊥ } ≤ (finrank R M : Cardinal.{w}) := by suffices Cardinal.lift.{v} #{ i // p i ≠ ⊥ } ≤ Cardinal.lift.{v} (finrank R M : Cardinal.{w}) by rwa [Cardinal.lift_le] at this @@ -276,8 +279,8 @@ theorem CompleteLattice.Independent.subtype_ne_bot_le_finrank_aux /-- If `p` is an independent family of submodules of a `R`-finite module `M`, then the number of nontrivial subspaces in the family `p` is finite. -/ -noncomputable def CompleteLattice.Independent.fintypeNeBotOfFiniteDimensional - {p : ι → Submodule R M} (hp : CompleteLattice.Independent p) : +noncomputable def iSupIndep.fintypeNeBotOfFiniteDimensional + {p : ι → Submodule R M} (hp : iSupIndep p) : Fintype { i : ι // p i ≠ ⊥ } := by suffices #{ i // p i ≠ ⊥ } < (ℵ₀ : Cardinal.{w}) by rw [Cardinal.lt_aleph0_iff_fintype] at this @@ -289,9 +292,9 @@ noncomputable def CompleteLattice.Independent.fintypeNeBotOfFiniteDimensional number of nontrivial subspaces in the family `p` is bounded above by the dimension of `M`. Note that the `Fintype` hypothesis required here can be provided by -`CompleteLattice.Independent.fintypeNeBotOfFiniteDimensional`. -/ -theorem CompleteLattice.Independent.subtype_ne_bot_le_finrank - {p : ι → Submodule R M} (hp : CompleteLattice.Independent p) [Fintype { i // p i ≠ ⊥ }] : +`iSupIndep.fintypeNeBotOfFiniteDimensional`. -/ +theorem iSupIndep.subtype_ne_bot_le_finrank + {p : ι → Submodule R M} (hp : iSupIndep p) [Fintype { i // p i ≠ ⊥ }] : Fintype.card { i // p i ≠ ⊥ } ≤ finrank R M := by simpa using hp.subtype_ne_bot_le_finrank_aux end diff --git a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean index 05c1324d366e9..3003ced8d18b7 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean @@ -664,11 +664,11 @@ lemma injOn_iSup_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) : apply injOn_maxGenEigenspace theorem independent_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) (k : ℕ∞) : - CompleteLattice.Independent (f.genEigenspace · k) := by + iSupIndep (f.genEigenspace · k) := by classical suffices ∀ μ₁ (s : Finset R), μ₁ ∉ s → Disjoint (f.genEigenspace μ₁ k) (s.sup fun μ ↦ f.genEigenspace μ k) by - simp_rw [CompleteLattice.independent_iff_supIndep_of_injOn (injOn_genEigenspace f k), + simp_rw [iSupIndep_iff_supIndep_of_injOn (injOn_genEigenspace f k), Finset.supIndep_iff_disjoint_erase] exact fun s μ _ ↦ this _ _ (s.not_mem_erase μ) intro μ₁ s @@ -703,27 +703,29 @@ theorem independent_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) (k : rwa [ih.eq_bot, Submodule.mem_bot] at hyz theorem independent_maxGenEigenspace [NoZeroSMulDivisors R M] (f : End R M) : - CompleteLattice.Independent f.maxGenEigenspace := by + iSupIndep f.maxGenEigenspace := by apply independent_genEigenspace @[deprecated independent_genEigenspace (since := "2024-10-23")] theorem independent_iSup_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) : - CompleteLattice.Independent (fun μ ↦ ⨆ k : ℕ, f.genEigenspace μ k) := by + iSupIndep (fun μ ↦ ⨆ k : ℕ, f.genEigenspace μ k) := by simp_rw [iSup_genEigenspace_eq] apply independent_maxGenEigenspace /-- The eigenspaces of a linear operator form an independent family of subspaces of `M`. That is, any eigenspace has trivial intersection with the span of all the other eigenspaces. -/ -theorem eigenspaces_independent [NoZeroSMulDivisors R M] (f : End R M) : - CompleteLattice.Independent f.eigenspace := +theorem eigenspaces_iSupIndep [NoZeroSMulDivisors R M] (f : End R M) : + iSupIndep f.eigenspace := (f.independent_genEigenspace 1).mono fun _ ↦ le_rfl +@[deprecated (since := "2024-11-24")] alias eigenspaces_independent := eigenspaces_iSupIndep + /-- Eigenvectors corresponding to distinct eigenvalues of a linear operator are linearly independent. -/ theorem eigenvectors_linearIndependent' {ι : Type*} [NoZeroSMulDivisors R M] (f : End R M) (μ : ι → R) (hμ : Function.Injective μ) (v : ι → M) (h_eigenvec : ∀ i, f.HasEigenvector (μ i) (v i)) : LinearIndependent R v := - f.eigenspaces_independent.comp hμ |>.linearIndependent _ + f.eigenspaces_iSupIndep.comp hμ |>.linearIndependent _ (fun i ↦ h_eigenvec i |>.left) (fun i ↦ h_eigenvec i |>.right) /-- Eigenvectors corresponding to distinct eigenvalues of a linear operator are linearly diff --git a/Mathlib/LinearAlgebra/Eigenspace/Matrix.lean b/Mathlib/LinearAlgebra/Eigenspace/Matrix.lean index 7ae63cc2e199b..0b327ae0fd03d 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Matrix.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Matrix.lean @@ -55,7 +55,7 @@ lemma hasEigenvalue_toLin_diagonal_iff (d : n → R) {μ : R} [NoZeroSMulDivisor rw [mem_eigenspace_iff] exact (hasEigenvector_toLin_diagonal d i b).apply_eq_smul have hμ_not_mem : μ ∉ Set.range d := by simpa using fun i ↦ (hμ i) - have := eigenspaces_independent (toLin b b (diagonal d)) |>.disjoint_biSup hμ_not_mem + have := eigenspaces_iSupIndep (toLin b b (diagonal d)) |>.disjoint_biSup hμ_not_mem rw [h_iSup, disjoint_top] at this exact h_eig this · rintro ⟨i, rfl⟩ diff --git a/Mathlib/LinearAlgebra/Eigenspace/Pi.lean b/Mathlib/LinearAlgebra/Eigenspace/Pi.lean index 9f089022d24ad..ddb8f9f9ea5ac 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Pi.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Pi.lean @@ -85,7 +85,7 @@ lemma injOn_iInf_maxGenEigenspace : lemma independent_iInf_maxGenEigenspace_of_forall_mapsTo (h : ∀ i j φ, MapsTo (f i) ((f j).maxGenEigenspace φ) ((f j).maxGenEigenspace φ)) : - CompleteLattice.Independent fun χ : ι → R ↦ ⨅ i, (f i).maxGenEigenspace (χ i) := by + iSupIndep fun χ : ι → R ↦ ⨅ i, (f i).maxGenEigenspace (χ i) := by replace h (l : ι) (χ : ι → R) : MapsTo (f l) (⨅ i, (f i).maxGenEigenspace (χ i)) (⨅ i, (f i).maxGenEigenspace (χ i)) := by intro x hx @@ -95,7 +95,7 @@ lemma independent_iInf_maxGenEigenspace_of_forall_mapsTo suffices ∀ χ (s : Finset (ι → R)) (_ : χ ∉ s), Disjoint (⨅ i, (f i).maxGenEigenspace (χ i)) (s.sup fun (χ : ι → R) ↦ ⨅ i, (f i).maxGenEigenspace (χ i)) by - simpa only [CompleteLattice.independent_iff_supIndep_of_injOn (injOn_iInf_maxGenEigenspace f), + simpa only [iSupIndep_iff_supIndep_of_injOn (injOn_iInf_maxGenEigenspace f), Finset.supIndep_iff_disjoint_erase] using fun s χ _ ↦ this _ _ (s.not_mem_erase χ) intro χ₁ s induction s using Finset.induction_on with diff --git a/Mathlib/LinearAlgebra/LinearIndependent.lean b/Mathlib/LinearAlgebra/LinearIndependent.lean index 39d2c59e03232..eedfde103f8a0 100644 --- a/Mathlib/LinearAlgebra/LinearIndependent.lean +++ b/Mathlib/LinearAlgebra/LinearIndependent.lean @@ -913,10 +913,10 @@ theorem linearIndependent_iff_not_smul_mem_span : simp [hij] · simp [hl]⟩ -/-- See also `CompleteLattice.independent_iff_linearIndependent_of_ne_zero`. -/ -theorem LinearIndependent.independent_span_singleton (hv : LinearIndependent R v) : - CompleteLattice.Independent fun i => R ∙ v i := by - refine CompleteLattice.independent_def.mp fun i => ?_ +/-- See also `iSupIndep_iff_linearIndependent_of_ne_zero`. -/ +theorem LinearIndependent.iSupIndep_span_singleton (hv : LinearIndependent R v) : + iSupIndep fun i => R ∙ v i := by + refine iSupIndep_def.mp fun i => ?_ rw [disjoint_iff_inf_le] intro m hm simp only [mem_inf, mem_span_singleton, iSup_subtype'] at hm @@ -928,6 +928,9 @@ theorem LinearIndependent.independent_span_singleton (hv : LinearIndependent R v ext simp +@[deprecated (since := "2024-11-24")] +alias LinearIndependent.independent_span_singleton := LinearIndependent.iSupIndep_span_singleton + variable (R) theorem exists_maximal_independent' (s : ι → M) : diff --git a/Mathlib/LinearAlgebra/Projectivization/Independence.lean b/Mathlib/LinearAlgebra/Projectivization/Independence.lean index e9a6bf429c7e2..403138f8a84d9 100644 --- a/Mathlib/LinearAlgebra/Projectivization/Independence.lean +++ b/Mathlib/LinearAlgebra/Projectivization/Independence.lean @@ -56,17 +56,19 @@ theorem independent_iff : Independent f ↔ LinearIndependent K (Projectivizatio /-- A family of points in projective space is independent if and only if the family of submodules which the points determine is independent in the lattice-theoretic sense. -/ -theorem independent_iff_completeLattice_independent : - Independent f ↔ CompleteLattice.Independent fun i => (f i).submodule := by +theorem independent_iff_iSupIndep : Independent f ↔ iSupIndep fun i => (f i).submodule := by refine ⟨?_, fun h => ?_⟩ · rintro ⟨f, hf, hi⟩ simp only [submodule_mk] - exact (CompleteLattice.independent_iff_linearIndependent_of_ne_zero (R := K) hf).mpr hi + exact (iSupIndep_iff_linearIndependent_of_ne_zero (R := K) hf).mpr hi · rw [independent_iff] refine h.linearIndependent (Projectivization.submodule ∘ f) (fun i => ?_) fun i => ?_ · simpa only [Function.comp_apply, submodule_eq] using Submodule.mem_span_singleton_self _ · exact rep_nonzero (f i) +@[deprecated (since := "2024-11-24")] +alias independent_iff_completeLattice_independent := independent_iff_iSupIndep + /-- A linearly dependent family of nonzero vectors gives a dependent family of points in projective space. -/ inductive Dependent : (ι → ℙ K V) → Prop diff --git a/Mathlib/Order/CompactlyGenerated/Basic.lean b/Mathlib/Order/CompactlyGenerated/Basic.lean index f3bc338a1d2ea..e802e01e3fb3e 100644 --- a/Mathlib/Order/CompactlyGenerated/Basic.lean +++ b/Mathlib/Order/CompactlyGenerated/Basic.lean @@ -273,13 +273,14 @@ alias ⟨_, IsSupClosedCompact.isSupFiniteCompact⟩ := isSupFiniteCompact_iff_i alias ⟨_, WellFoundedGT.isSupClosedCompact⟩ := isSupClosedCompact_iff_wellFoundedGT -variable {α} +end CompleteLattice + -theorem WellFoundedGT.finite_of_setIndependent [WellFoundedGT α] {s : Set α} - (hs : SetIndependent s) : s.Finite := by +theorem WellFoundedGT.finite_of_sSupIndep [WellFoundedGT α] {s : Set α} + (hs : sSupIndep s) : s.Finite := by classical refine Set.not_infinite.mp fun contra => ?_ - obtain ⟨t, ht₁, ht₂⟩ := WellFoundedGT.isSupFiniteCompact α s + obtain ⟨t, ht₁, ht₂⟩ := CompleteLattice.WellFoundedGT.isSupFiniteCompact α s replace contra : ∃ x : α, x ∈ s ∧ x ≠ ⊥ ∧ x ∉ t := by have : (s \ (insert ⊥ t : Finset α)).Infinite := contra.diff (Finset.finite_toSet _) obtain ⟨x, hx₁, hx₂⟩ := this.nonempty @@ -290,41 +291,59 @@ theorem WellFoundedGT.finite_of_setIndependent [WellFoundedGT α] {s : Set α} simpa [Disjoint, hx₂, ← t.sup_id_eq_sSup, ← ht₂] using this.eq_bot apply hx₁ rw [← hs, eq_comm, inf_eq_left] - exact le_sSup _ _ hx₀ + exact le_sSup hx₀ -theorem WellFoundedGT.finite_ne_bot_of_independent [WellFoundedGT α] - {ι : Type*} {t : ι → α} (ht : Independent t) : Set.Finite {i | t i ≠ ⊥} := by +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.WellFoundedGT.finite_of_setIndependent := WellFoundedGT.finite_of_sSupIndep + +theorem WellFoundedGT.finite_ne_bot_of_iSupIndep [WellFoundedGT α] + {ι : Type*} {t : ι → α} (ht : iSupIndep t) : Set.Finite {i | t i ≠ ⊥} := by refine Finite.of_finite_image (Finite.subset ?_ (image_subset_range t _)) ht.injOn - exact WellFoundedGT.finite_of_setIndependent ht.setIndependent_range + exact WellFoundedGT.finite_of_sSupIndep ht.sSupIndep_range + +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.WellFoundedGT.finite_ne_bot_of_independent := + WellFoundedGT.finite_ne_bot_of_iSupIndep -theorem WellFoundedGT.finite_of_independent [WellFoundedGT α] {ι : Type*} - {t : ι → α} (ht : Independent t) (h_ne_bot : ∀ i, t i ≠ ⊥) : Finite ι := - haveI := (WellFoundedGT.finite_of_setIndependent ht.setIndependent_range).to_subtype +theorem WellFoundedGT.finite_of_iSupIndep [WellFoundedGT α] {ι : Type*} + {t : ι → α} (ht : iSupIndep t) (h_ne_bot : ∀ i, t i ≠ ⊥) : Finite ι := + haveI := (WellFoundedGT.finite_of_sSupIndep ht.sSupIndep_range).to_subtype Finite.of_injective_finite_range (ht.injective h_ne_bot) -theorem WellFoundedLT.finite_of_setIndependent [WellFoundedLT α] {s : Set α} - (hs : SetIndependent s) : s.Finite := by +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.WellFoundedGT.finite_of_independent := WellFoundedGT.finite_of_iSupIndep + +theorem WellFoundedLT.finite_of_sSupIndep [WellFoundedLT α] {s : Set α} + (hs : sSupIndep s) : s.Finite := by by_contra inf let e := (Infinite.diff inf <| finite_singleton ⊥).to_subtype.natEmbedding let a n := ⨆ i ≥ n, (e i).1 have sup_le n : (e n).1 ⊔ a (n + 1) ≤ a n := sup_le_iff.mpr ⟨le_iSup₂_of_le n le_rfl le_rfl, iSup₂_le fun i hi ↦ le_iSup₂_of_le i (n.le_succ.trans hi) le_rfl⟩ have lt n : a (n + 1) < a n := (Disjoint.right_lt_sup_of_left_ne_bot - ((hs (e n).2.1).mono_right <| iSup₂_le fun i hi ↦ le_sSup _ _ ?_) (e n).2.2).trans_le (sup_le n) + ((hs (e n).2.1).mono_right <| iSup₂_le fun i hi ↦ le_sSup ?_) (e n).2.2).trans_le (sup_le n) · exact (RelEmbedding.natGT a lt).not_wellFounded_of_decreasing_seq wellFounded_lt exact ⟨(e i).2.1, fun h ↦ n.lt_succ_self.not_le <| hi.trans_eq <| e.2 <| Subtype.val_injective h⟩ -theorem WellFoundedLT.finite_ne_bot_of_independent [WellFoundedLT α] - {ι : Type*} {t : ι → α} (ht : Independent t) : Set.Finite {i | t i ≠ ⊥} := by +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.WellFoundedLT.finite_of_setIndependent := WellFoundedLT.finite_of_sSupIndep + +theorem WellFoundedLT.finite_ne_bot_of_iSupIndep [WellFoundedLT α] + {ι : Type*} {t : ι → α} (ht : iSupIndep t) : Set.Finite {i | t i ≠ ⊥} := by refine Finite.of_finite_image (Finite.subset ?_ (image_subset_range t _)) ht.injOn - exact WellFoundedLT.finite_of_setIndependent ht.setIndependent_range + exact WellFoundedLT.finite_of_sSupIndep ht.sSupIndep_range + +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.WellFoundedLT.finite_ne_bot_of_independent := + WellFoundedLT.finite_ne_bot_of_iSupIndep -theorem WellFoundedLT.finite_of_independent [WellFoundedLT α] {ι : Type*} - {t : ι → α} (ht : Independent t) (h_ne_bot : ∀ i, t i ≠ ⊥) : Finite ι := - haveI := (WellFoundedLT.finite_of_setIndependent ht.setIndependent_range).to_subtype +theorem WellFoundedLT.finite_of_iSupIndep [WellFoundedLT α] {ι : Type*} + {t : ι → α} (ht : iSupIndep t) (h_ne_bot : ∀ i, t i ≠ ⊥) : Finite ι := + haveI := (WellFoundedLT.finite_of_sSupIndep ht.sSupIndep_range).to_subtype Finite.of_injective_finite_range (ht.injective h_ne_bot) -end CompleteLattice +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.WellFoundedLT.finite_of_independent := WellFoundedLT.finite_of_iSupIndep /-- A complete lattice is said to be compactly generated if any element is the `sSup` of compact elements. -/ @@ -413,9 +432,9 @@ theorem inf_sSup_eq_iSup_inf_sup_finset : (iSup_le fun t => iSup_le fun h => inf_le_inf_left _ ((Finset.sup_id_eq_sSup t).symm ▸ sSup_le_sSup h)) -theorem CompleteLattice.setIndependent_iff_finite {s : Set α} : - CompleteLattice.SetIndependent s ↔ - ∀ t : Finset α, ↑t ⊆ s → CompleteLattice.SetIndependent (↑t : Set α) := +theorem sSupIndep_iff_finite {s : Set α} : + sSupIndep s ↔ + ∀ t : Finset α, ↑t ⊆ s → sSupIndep (↑t : Set α) := ⟨fun hs _ ht => hs.mono ht, fun h a ha => by rw [disjoint_iff, inf_sSup_eq_iSup_inf_sup_finset, iSup_eq_bot] intro t @@ -428,10 +447,13 @@ theorem CompleteLattice.setIndependent_iff_finite {s : Set α} : · rw [Finset.coe_insert, Set.insert_subset_iff] exact ⟨ha, Set.Subset.trans ht diff_subset⟩⟩ -lemma CompleteLattice.independent_iff_supIndep_of_injOn {ι : Type*} {f : ι → α} +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.setIndependent_iff_finite := sSupIndep_iff_finite + +lemma iSupIndep_iff_supIndep_of_injOn {ι : Type*} {f : ι → α} (hf : InjOn f {i | f i ≠ ⊥}) : - CompleteLattice.Independent f ↔ ∀ (s : Finset ι), s.SupIndep f := by - refine ⟨fun h ↦ h.supIndep', fun h ↦ CompleteLattice.independent_def'.mpr fun i ↦ ?_⟩ + iSupIndep f ↔ ∀ (s : Finset ι), s.SupIndep f := by + refine ⟨fun h ↦ h.supIndep', fun h ↦ iSupIndep_def'.mpr fun i ↦ ?_⟩ simp_rw [disjoint_iff, inf_sSup_eq_iSup_inf_sup_finset, iSup_eq_bot, ← disjoint_iff] intro s hs classical @@ -451,11 +473,14 @@ lemma CompleteLattice.independent_iff_supIndep_of_injOn {ι : Type*} {f : ι → rw [Finset.supIndep_iff_disjoint_erase] at h exact h i (Finset.mem_insert_self i _) -theorem CompleteLattice.setIndependent_iUnion_of_directed {η : Type*} {s : η → Set α} - (hs : Directed (· ⊆ ·) s) (h : ∀ i, CompleteLattice.SetIndependent (s i)) : - CompleteLattice.SetIndependent (⋃ i, s i) := by +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.independent_iff_supIndep_of_injOn := iSupIndep_iff_supIndep_of_injOn + +theorem sSupIndep_iUnion_of_directed {η : Type*} {s : η → Set α} + (hs : Directed (· ⊆ ·) s) (h : ∀ i, sSupIndep (s i)) : + sSupIndep (⋃ i, s i) := by by_cases hη : Nonempty η - · rw [CompleteLattice.setIndependent_iff_finite] + · rw [sSupIndep_iff_finite] intro t ht obtain ⟨I, fi, hI⟩ := Set.finite_subset_iUnion t.finite_toSet ht obtain ⟨i, hi⟩ := hs.finset_le fi.toFinset @@ -465,10 +490,16 @@ theorem CompleteLattice.setIndependent_iUnion_of_directed {η : Type*} {s : η exfalso exact hη ⟨i⟩ -theorem CompleteLattice.independent_sUnion_of_directed {s : Set (Set α)} (hs : DirectedOn (· ⊆ ·) s) - (h : ∀ a ∈ s, CompleteLattice.SetIndependent a) : CompleteLattice.SetIndependent (⋃₀ s) := by +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.setIndependent_iUnion_of_directed := sSupIndep_iUnion_of_directed + +theorem iSupIndep_sUnion_of_directed {s : Set (Set α)} (hs : DirectedOn (· ⊆ ·) s) + (h : ∀ a ∈ s, sSupIndep a) : sSupIndep (⋃₀ s) := by rw [Set.sUnion_eq_iUnion] - exact CompleteLattice.setIndependent_iUnion_of_directed hs.directed_val (by simpa using h) + exact sSupIndep_iUnion_of_directed hs.directed_val (by simpa using h) + +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.independent_sUnion_of_directed := iSupIndep_sUnion_of_directed end @@ -495,11 +526,11 @@ alias IsSupFiniteCompact.wellFounded := IsSupFiniteCompact.wellFoundedGT @[deprecated (since := "2024-10-07")] alias _root_.WellFounded.isSupClosedCompact := WellFoundedGT.isSupClosedCompact @[deprecated (since := "2024-10-07")] -alias WellFounded.finite_of_setIndependent := WellFoundedGT.finite_of_setIndependent +alias WellFounded.finite_of_setIndependent := WellFoundedGT.finite_of_sSupIndep @[deprecated (since := "2024-10-07")] -alias WellFounded.finite_ne_bot_of_independent := WellFoundedGT.finite_ne_bot_of_independent +alias WellFounded.finite_ne_bot_of_independent := WellFoundedGT.finite_ne_bot_of_iSupIndep @[deprecated (since := "2024-10-07")] -alias WellFounded.finite_of_independent := WellFoundedGT.finite_of_independent +alias WellFounded.finite_of_independent := WellFoundedGT.finite_of_iSupIndep @[deprecated (since := "2024-10-07")] alias isCompactlyGenerated_of_wellFounded := isCompactlyGenerated_of_wellFoundedGT @@ -578,16 +609,16 @@ Most explicitly, every element is the complement of a supremum of indepedendent /-- In an atomic lattice, every element `b` has a complement of the form `sSup s`, where each element of `s` is an atom. See also `complementedLattice_of_sSup_atoms_eq_top`. -/ -theorem exists_setIndependent_isCompl_sSup_atoms (h : sSup { a : α | IsAtom a } = ⊤) (b : α) : - ∃ s : Set α, CompleteLattice.SetIndependent s ∧ +theorem exists_sSupIndep_isCompl_sSup_atoms (h : sSup { a : α | IsAtom a } = ⊤) (b : α) : + ∃ s : Set α, sSupIndep s ∧ IsCompl b (sSup s) ∧ ∀ ⦃a⦄, a ∈ s → IsAtom a := by -- porting note(https://github.com/leanprover-community/mathlib4/issues/5732): -- `obtain` chokes on the placeholder. have zorn := zorn_subset - (S := {s : Set α | CompleteLattice.SetIndependent s ∧ Disjoint b (sSup s) ∧ ∀ a ∈ s, IsAtom a}) + (S := {s : Set α | sSupIndep s ∧ Disjoint b (sSup s) ∧ ∀ a ∈ s, IsAtom a}) fun c hc1 hc2 => ⟨⋃₀ c, - ⟨CompleteLattice.independent_sUnion_of_directed hc2.directedOn fun s hs => (hc1 hs).1, ?_, + ⟨iSupIndep_sUnion_of_directed hc2.directedOn fun s hs => (hc1 hs).1, ?_, fun a ⟨s, sc, as⟩ => (hc1 sc).2.2 a as⟩, fun _ => Set.subset_sUnion_of_mem⟩ swap @@ -633,16 +664,22 @@ theorem exists_setIndependent_isCompl_sSup_atoms (h : sSup { a : α | IsAtom a } · exact s_atoms x hx · exact ha -theorem exists_setIndependent_of_sSup_atoms_eq_top (h : sSup { a : α | IsAtom a } = ⊤) : - ∃ s : Set α, CompleteLattice.SetIndependent s ∧ sSup s = ⊤ ∧ ∀ ⦃a⦄, a ∈ s → IsAtom a := - let ⟨s, s_ind, s_top, s_atoms⟩ := exists_setIndependent_isCompl_sSup_atoms h ⊥ +@[deprecated (since := "2024-11-24")] +alias exists_setIndependent_isCompl_sSup_atoms := exists_sSupIndep_isCompl_sSup_atoms + +theorem exists_sSupIndep_of_sSup_atoms_eq_top (h : sSup { a : α | IsAtom a } = ⊤) : + ∃ s : Set α, sSupIndep s ∧ sSup s = ⊤ ∧ ∀ ⦃a⦄, a ∈ s → IsAtom a := + let ⟨s, s_ind, s_top, s_atoms⟩ := exists_sSupIndep_isCompl_sSup_atoms h ⊥ ⟨s, s_ind, eq_top_of_isCompl_bot s_top.symm, s_atoms⟩ +@[deprecated (since := "2024-11-24")] +alias exists_setIndependent_of_sSup_atoms_eq_top := exists_sSupIndep_of_sSup_atoms_eq_top + /-- See [Theorem 6.6][calugareanu]. -/ theorem complementedLattice_of_sSup_atoms_eq_top (h : sSup { a : α | IsAtom a } = ⊤) : ComplementedLattice α := ⟨fun b => - let ⟨s, _, s_top, _⟩ := exists_setIndependent_isCompl_sSup_atoms h b + let ⟨s, _, s_top, _⟩ := exists_sSupIndep_isCompl_sSup_atoms h b ⟨sSup s, s_top⟩⟩ /-- See [Theorem 6.6][calugareanu]. -/ diff --git a/Mathlib/Order/SupIndep.lean b/Mathlib/Order/SupIndep.lean index 68e516ec2adec..27770ea7eaa4a 100644 --- a/Mathlib/Order/SupIndep.lean +++ b/Mathlib/Order/SupIndep.lean @@ -18,19 +18,19 @@ sup-independent if, for all `a`, `f a` and the supremum of the rest are disjoint ## Main definitions * `Finset.SupIndep s f`: a family of elements `f` are supremum independent on the finite set `s`. -* `CompleteLattice.SetIndependent s`: a set of elements are supremum independent. -* `CompleteLattice.Independent f`: a family of elements are supremum independent. +* `sSupIndep s`: a set of elements are supremum independent. +* `iSupIndep f`: a family of elements are supremum independent. ## Main statements * In a distributive lattice, supremum independence is equivalent to pairwise disjointness: * `Finset.supIndep_iff_pairwiseDisjoint` - * `CompleteLattice.setIndependent_iff_pairwiseDisjoint` - * `CompleteLattice.independent_iff_pairwiseDisjoint` + * `CompleteLattice.sSupIndep_iff_pairwiseDisjoint` + * `CompleteLattice.iSupIndep_iff_pairwiseDisjoint` * Otherwise, supremum independence is stronger than pairwise disjointness: * `Finset.SupIndep.pairwiseDisjoint` - * `CompleteLattice.SetIndependent.pairwiseDisjoint` - * `CompleteLattice.Independent.pairwiseDisjoint` + * `sSupIndep.pairwiseDisjoint` + * `iSupIndep.pairwiseDisjoint` ## Implementation notes @@ -88,7 +88,7 @@ theorem SupIndep.le_sup_iff (hs : s.SupIndep f) (hts : t ⊆ s) (hi : i ∈ s) ( by_contra hit exact hf i (disjoint_self.1 <| (hs hts hi hit).mono_right h) -/-- The RHS looks like the definition of `CompleteLattice.Independent`. -/ +/-- The RHS looks like the definition of `iSupIndep`. -/ theorem supIndep_iff_disjoint_erase [DecidableEq ι] : s.SupIndep f ↔ ∀ i ∈ s, Disjoint (f i) ((s.erase i).sup f) := ⟨fun hs _ hi => hs (erase_subset _ _) hi (not_mem_erase _ _), fun hs _ ht i hi hit => @@ -269,38 +269,48 @@ end Finset /-! ### On complete lattices via `sSup` -/ - -namespace CompleteLattice - +section CompleteLattice variable [CompleteLattice α] open Set Function /-- An independent set of elements in a complete lattice is one in which every element is disjoint from the `Sup` of the rest. -/ -def SetIndependent (s : Set α) : Prop := +def sSupIndep (s : Set α) : Prop := ∀ ⦃a⦄, a ∈ s → Disjoint a (sSup (s \ {a})) -variable {s : Set α} (hs : SetIndependent s) +@[deprecated (since := "2024-11-24")] alias CompleteLattice.SetIndependent := sSupIndep + +variable {s : Set α} (hs : sSupIndep s) @[simp] -theorem setIndependent_empty : SetIndependent (∅ : Set α) := fun x hx => +theorem sSupIndep_empty : sSupIndep (∅ : Set α) := fun x hx => (Set.not_mem_empty x hx).elim +@[deprecated (since := "2024-11-24")] alias CompleteLattice.setIndependent_empty := sSupIndep_empty + include hs in -theorem SetIndependent.mono {t : Set α} (hst : t ⊆ s) : SetIndependent t := fun _ ha => +theorem sSupIndep.mono {t : Set α} (hst : t ⊆ s) : sSupIndep t := fun _ ha => (hs (hst ha)).mono_right (sSup_le_sSup (diff_subset_diff_left hst)) +@[deprecated (since := "2024-11-24")] alias CompleteLattice.SetIndependent.mono := sSupIndep.mono + include hs in /-- If the elements of a set are independent, then any pair within that set is disjoint. -/ -theorem SetIndependent.pairwiseDisjoint : s.PairwiseDisjoint id := fun _ hx y hy h => +theorem sSupIndep.pairwiseDisjoint : s.PairwiseDisjoint id := fun _ hx y hy h => disjoint_sSup_right (hs hx) ((mem_diff y).mpr ⟨hy, h.symm⟩) -theorem setIndependent_singleton (a : α) : SetIndependent ({a} : Set α) := fun i hi ↦ by +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.SetIndependent.pairwiseDisjoint := sSupIndep.pairwiseDisjoint + +theorem sSupIndep_singleton (a : α) : sSupIndep ({a} : Set α) := fun i hi ↦ by simp_all -theorem setIndependent_pair {a b : α} (hab : a ≠ b) : - SetIndependent ({a, b} : Set α) ↔ Disjoint a b := by +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.setIndependent_singleton := sSupIndep_singleton + +theorem sSupIndep_pair {a b : α} (hab : a ≠ b) : + sSupIndep ({a, b} : Set α) ↔ Disjoint a b := by constructor · intro h exact h.pairwiseDisjoint (mem_insert _ _) (mem_insert_of_mem _ (mem_singleton _)) hab @@ -310,15 +320,20 @@ theorem setIndependent_pair {a b : α} (hab : a ≠ b) : · convert h.symm using 1 simp [hab, sSup_singleton] +@[deprecated (since := "2024-11-24")] alias CompleteLattice.setIndependent_pair := sSupIndep_pair + include hs in /-- If the elements of a set are independent, then any element is disjoint from the `sSup` of some subset of the rest. -/ -theorem SetIndependent.disjoint_sSup {x : α} {y : Set α} (hx : x ∈ s) (hy : y ⊆ s) (hxy : x ∉ y) : +theorem sSupIndep.disjoint_sSup {x : α} {y : Set α} (hx : x ∈ s) (hy : y ⊆ s) (hxy : x ∉ y) : Disjoint x (sSup y) := by have := (hs.mono <| insert_subset_iff.mpr ⟨hx, hy⟩) (mem_insert x _) rw [insert_diff_of_mem _ (mem_singleton _), diff_singleton_eq_self hxy] at this exact this +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.SetIndependent.disjoint_sSup := sSupIndep.disjoint_sSup + /-- An independent indexed family of elements in a complete lattice is one in which every element is disjoint from the `iSup` of the rest. @@ -328,70 +343,96 @@ theorem SetIndependent.disjoint_sSup {x : α} {y : Set α} (hx : x ∈ s) (hy : Example: an indexed family of submodules of a module is independent in this sense if and only the natural map from the direct sum of the submodules to the module is injective. -/ -def Independent {ι : Sort*} {α : Type*} [CompleteLattice α] (t : ι → α) : Prop := +def iSupIndep {ι : Sort*} {α : Type*} [CompleteLattice α] (t : ι → α) : Prop := ∀ i : ι, Disjoint (t i) (⨆ (j) (_ : j ≠ i), t j) -theorem setIndependent_iff {α : Type*} [CompleteLattice α] (s : Set α) : - SetIndependent s ↔ Independent ((↑) : s → α) := by - simp_rw [Independent, SetIndependent, SetCoe.forall, sSup_eq_iSup] +@[deprecated (since := "2024-11-24")] alias CompleteLattice.Independent := iSupIndep + +theorem sSupIndep_iff {α : Type*} [CompleteLattice α] (s : Set α) : + sSupIndep s ↔ iSupIndep ((↑) : s → α) := by + simp_rw [iSupIndep, sSupIndep, SetCoe.forall, sSup_eq_iSup] refine forall₂_congr fun a ha => ?_ simp [iSup_subtype, iSup_and] -variable {t : ι → α} (ht : Independent t) +@[deprecated (since := "2024-11-24")] alias CompleteLattice.setIndependent_iff := sSupIndep_iff + +variable {t : ι → α} (ht : iSupIndep t) -theorem independent_def : Independent t ↔ ∀ i : ι, Disjoint (t i) (⨆ (j) (_ : j ≠ i), t j) := +theorem iSupIndep_def : iSupIndep t ↔ ∀ i, Disjoint (t i) (⨆ (j) (_ : j ≠ i), t j) := Iff.rfl -theorem independent_def' : Independent t ↔ ∀ i, Disjoint (t i) (sSup (t '' { j | j ≠ i })) := by +@[deprecated (since := "2024-11-24")] alias CompleteLattice.independent_def := iSupIndep_def + +theorem iSupIndep_def' : iSupIndep t ↔ ∀ i, Disjoint (t i) (sSup (t '' { j | j ≠ i })) := by simp_rw [sSup_image] rfl -theorem independent_def'' : - Independent t ↔ ∀ i, Disjoint (t i) (sSup { a | ∃ j ≠ i, t j = a }) := by - rw [independent_def'] +@[deprecated (since := "2024-11-24")] alias CompleteLattice.independent_def' := iSupIndep_def' + +theorem iSupIndep_def'' : + iSupIndep t ↔ ∀ i, Disjoint (t i) (sSup { a | ∃ j ≠ i, t j = a }) := by + rw [iSupIndep_def'] aesop +@[deprecated (since := "2024-11-24")] alias CompleteLattice.independent_def'' := iSupIndep_def'' + @[simp] -theorem independent_empty (t : Empty → α) : Independent t := +theorem iSupIndep_empty (t : Empty → α) : iSupIndep t := nofun +@[deprecated (since := "2024-11-24")] alias CompleteLattice.independent_empty := iSupIndep_empty + @[simp] -theorem independent_pempty (t : PEmpty → α) : Independent t := +theorem iSupIndep_pempty (t : PEmpty → α) : iSupIndep t := nofun +@[deprecated (since := "2024-11-24")] alias CompleteLattice.independent_pempty := iSupIndep_pempty + include ht in /-- If the elements of a set are independent, then any pair within that set is disjoint. -/ -theorem Independent.pairwiseDisjoint : Pairwise (Disjoint on t) := fun x y h => +theorem iSupIndep.pairwiseDisjoint : Pairwise (Disjoint on t) := fun x y h => disjoint_sSup_right (ht x) ⟨y, iSup_pos h.symm⟩ -theorem Independent.mono {s t : ι → α} (hs : Independent s) (hst : t ≤ s) : Independent t := +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.Independent.pairwiseDisjoint := iSupIndep.pairwiseDisjoint + +theorem iSupIndep.mono {s t : ι → α} (hs : iSupIndep s) (hst : t ≤ s) : iSupIndep t := fun i => (hs i).mono (hst i) <| iSup₂_mono fun j _ => hst j +@[deprecated (since := "2024-11-24")] alias CompleteLattice.Independent.mono := iSupIndep.mono + /-- Composing an independent indexed family with an injective function on the index results in another indepedendent indexed family. -/ -theorem Independent.comp {ι ι' : Sort*} {t : ι → α} {f : ι' → ι} (ht : Independent t) - (hf : Injective f) : Independent (t ∘ f) := fun i => +theorem iSupIndep.comp {ι ι' : Sort*} {t : ι → α} {f : ι' → ι} (ht : iSupIndep t) + (hf : Injective f) : iSupIndep (t ∘ f) := fun i => (ht (f i)).mono_right <| by refine (iSup_mono fun i => ?_).trans (iSup_comp_le _ f) exact iSup_const_mono hf.ne -theorem Independent.comp' {ι ι' : Sort*} {t : ι → α} {f : ι' → ι} (ht : Independent <| t ∘ f) - (hf : Surjective f) : Independent t := by +@[deprecated (since := "2024-11-24")] alias CompleteLattice.Independent.comp := iSupIndep.comp + +theorem iSupIndep.comp' {ι ι' : Sort*} {t : ι → α} {f : ι' → ι} (ht : iSupIndep <| t ∘ f) + (hf : Surjective f) : iSupIndep t := by intro i obtain ⟨i', rfl⟩ := hf i rw [← hf.iSup_comp] exact (ht i').mono_right (biSup_mono fun j' hij => mt (congr_arg f) hij) -theorem Independent.setIndependent_range (ht : Independent t) : SetIndependent <| range t := by - rw [setIndependent_iff] +@[deprecated (since := "2024-11-24")] alias CompleteLattice.Independent.comp' := iSupIndep.comp' + +theorem iSupIndep.sSupIndep_range (ht : iSupIndep t) : sSupIndep <| range t := by + rw [sSupIndep_iff] rw [← coe_comp_rangeFactorization t] at ht exact ht.comp' surjective_onto_range +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.Independent.setIndependent_range := iSupIndep.sSupIndep_range + @[simp] -theorem independent_ne_bot_iff_independent : - Independent (fun i : {i // t i ≠ ⊥} ↦ t i) ↔ Independent t := by +theorem iSupIndep_ne_bot : + iSupIndep (fun i : {i // t i ≠ ⊥} ↦ t i) ↔ iSupIndep t := by refine ⟨fun h ↦ ?_, fun h ↦ h.comp Subtype.val_injective⟩ - simp only [independent_def] at h ⊢ + simp only [iSupIndep_def] at h ⊢ intro i cases eq_or_ne (t i) ⊥ with | inl hi => simp [hi] @@ -402,7 +443,10 @@ theorem independent_ne_bot_iff_independent : simp only [iSup_comm (ι' := _ ≠ i), this, ne_eq, sup_of_le_right, Subtype.mk.injEq, iSup_bot, bot_le] -theorem Independent.injOn (ht : Independent t) : InjOn t {i | t i ≠ ⊥} := by +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.independent_ne_bot_iff_independent := iSupIndep_ne_bot + +theorem iSupIndep.injOn (ht : iSupIndep t) : InjOn t {i | t i ≠ ⊥} := by rintro i _ j (hj : t j ≠ ⊥) h by_contra! contra apply hj @@ -413,12 +457,17 @@ theorem Independent.injOn (ht : Independent t) : InjOn t {i | t i ≠ ⊥} := by -- Porting note: needs explicit `f` exact le_iSup₂ (f := fun x _ ↦ t x) j contra -theorem Independent.injective (ht : Independent t) (h_ne_bot : ∀ i, t i ≠ ⊥) : Injective t := by +@[deprecated (since := "2024-11-24")] alias CompleteLattice.Independent.injOn := iSupIndep.injOn + +theorem iSupIndep.injective (ht : iSupIndep t) (h_ne_bot : ∀ i, t i ≠ ⊥) : Injective t := by suffices univ = {i | t i ≠ ⊥} by rw [injective_iff_injOn_univ, this]; exact ht.injOn aesop -theorem independent_pair {i j : ι} (hij : i ≠ j) (huniv : ∀ k, k = i ∨ k = j) : - Independent t ↔ Disjoint (t i) (t j) := by +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.Independent.injective := iSupIndep.injective + +theorem iSupIndep_pair {i j : ι} (hij : i ≠ j) (huniv : ∀ k, k = i ∨ k = j) : + iSupIndep t ↔ Disjoint (t i) (t j) := by constructor · exact fun h => h.pairwiseDisjoint hij · rintro h k @@ -428,37 +477,49 @@ theorem independent_pair {i j : ι} (hij : i ≠ j) (huniv : ∀ k, k = i ∨ k · refine h.symm.mono_right (iSup_le fun j => iSup_le fun hj => Eq.le ?_) rw [(huniv j).resolve_right hj] +@[deprecated (since := "2024-11-24")] alias CompleteLattice.independent_pair := iSupIndep_pair + /-- Composing an independent indexed family with an order isomorphism on the elements results in another independent indexed family. -/ -theorem Independent.map_orderIso {ι : Sort*} {α β : Type*} [CompleteLattice α] - [CompleteLattice β] (f : α ≃o β) {a : ι → α} (ha : Independent a) : Independent (f ∘ a) := +theorem iSupIndep.map_orderIso {ι : Sort*} {α β : Type*} [CompleteLattice α] + [CompleteLattice β] (f : α ≃o β) {a : ι → α} (ha : iSupIndep a) : iSupIndep (f ∘ a) := fun i => ((ha i).map_orderIso f).mono_right (f.monotone.le_map_iSup₂ _) +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.Independent.map_orderIso := iSupIndep.map_orderIso + @[simp] -theorem independent_map_orderIso_iff {ι : Sort*} {α β : Type*} [CompleteLattice α] - [CompleteLattice β] (f : α ≃o β) {a : ι → α} : Independent (f ∘ a) ↔ Independent a := +theorem iSupIndep_map_orderIso_iff {ι : Sort*} {α β : Type*} [CompleteLattice α] + [CompleteLattice β] (f : α ≃o β) {a : ι → α} : iSupIndep (f ∘ a) ↔ iSupIndep a := ⟨fun h => have hf : f.symm ∘ f ∘ a = a := congr_arg (· ∘ a) f.left_inv.comp_eq_id hf ▸ h.map_orderIso f.symm, fun h => h.map_orderIso f⟩ +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.independent_map_orderIso_iff := iSupIndep_map_orderIso_iff + /-- If the elements of a set are independent, then any element is disjoint from the `iSup` of some subset of the rest. -/ -theorem Independent.disjoint_biSup {ι : Type*} {α : Type*} [CompleteLattice α] {t : ι → α} - (ht : Independent t) {x : ι} {y : Set ι} (hx : x ∉ y) : Disjoint (t x) (⨆ i ∈ y, t i) := +theorem iSupIndep.disjoint_biSup {ι : Type*} {α : Type*} [CompleteLattice α] {t : ι → α} + (ht : iSupIndep t) {x : ι} {y : Set ι} (hx : x ∉ y) : Disjoint (t x) (⨆ i ∈ y, t i) := Disjoint.mono_right (biSup_mono fun _ hi => (ne_of_mem_of_not_mem hi hx : _)) (ht x) -lemma independent_of_independent_coe_Iic_comp {ι : Sort*} {a : α} {t : ι → Set.Iic a} - (ht : Independent ((↑) ∘ t : ι → α)) : Independent t := by +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.Independent.disjoint_biSup := iSupIndep.disjoint_biSup + +lemma iSupIndep.of_coe_Iic_comp {ι : Sort*} {a : α} {t : ι → Set.Iic a} + (ht : iSupIndep ((↑) ∘ t : ι → α)) : iSupIndep t := by intro i x specialize ht i simp_rw [Function.comp_apply, ← Set.Iic.coe_iSup] at ht exact @ht x -end CompleteLattice +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.independent_of_independent_coe_Iic_comp := iSupIndep.of_coe_Iic_comp -theorem CompleteLattice.independent_iff_supIndep [CompleteLattice α] {s : Finset ι} {f : ι → α} : - CompleteLattice.Independent (f ∘ ((↑) : s → ι)) ↔ s.SupIndep f := by +theorem iSupIndep_iff_supIndep {s : Finset ι} {f : ι → α} : + iSupIndep (f ∘ ((↑) : s → ι)) ↔ s.SupIndep f := by classical rw [Finset.supIndep_iff_disjoint_erase] refine Subtype.forall.trans (forall₂_congr fun a b => ?_) @@ -468,39 +529,47 @@ theorem CompleteLattice.independent_iff_supIndep [CompleteLattice α] {s : Finse congr! 1 simp [iSup_and, @iSup_comm _ (_ ∈ s)] -alias ⟨CompleteLattice.Independent.supIndep, Finset.SupIndep.independent⟩ := - CompleteLattice.independent_iff_supIndep +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.independent_iff_supIndep := iSupIndep_iff_supIndep + +alias ⟨iSupIndep.supIndep, Finset.SupIndep.independent⟩ := iSupIndep_iff_supIndep -theorem CompleteLattice.Independent.supIndep' [CompleteLattice α] {f : ι → α} (s : Finset ι) - (h : CompleteLattice.Independent f) : s.SupIndep f := - CompleteLattice.Independent.supIndep (h.comp Subtype.coe_injective) +theorem iSupIndep.supIndep' {f : ι → α} (s : Finset ι) (h : iSupIndep f) : s.SupIndep f := + iSupIndep.supIndep (h.comp Subtype.coe_injective) -/-- A variant of `CompleteLattice.independent_iff_supIndep` for `Fintype`s. -/ -theorem CompleteLattice.independent_iff_supIndep_univ [CompleteLattice α] [Fintype ι] {f : ι → α} : - CompleteLattice.Independent f ↔ Finset.univ.SupIndep f := by +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.Independent.supIndep' := iSupIndep.supIndep' + +/-- A variant of `CompleteLattice.iSupIndep_iff_supIndep` for `Fintype`s. -/ +theorem iSupIndep_iff_supIndep_univ [Fintype ι] {f : ι → α} : + iSupIndep f ↔ Finset.univ.SupIndep f := by classical - simp [Finset.supIndep_iff_disjoint_erase, CompleteLattice.Independent, Finset.sup_eq_iSup] + simp [Finset.supIndep_iff_disjoint_erase, iSupIndep, Finset.sup_eq_iSup] -alias ⟨CompleteLattice.Independent.sup_indep_univ, Finset.SupIndep.independent_of_univ⟩ := - CompleteLattice.independent_iff_supIndep_univ +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.independent_iff_supIndep_univ := iSupIndep_iff_supIndep_univ -section Frame +alias ⟨iSupIndep.sup_indep_univ, Finset.SupIndep.iSupIndep_of_univ⟩ := iSupIndep_iff_supIndep_univ -namespace CompleteLattice +end CompleteLattice +section Frame variable [Order.Frame α] -theorem setIndependent_iff_pairwiseDisjoint {s : Set α} : - SetIndependent s ↔ s.PairwiseDisjoint id := - ⟨SetIndependent.pairwiseDisjoint, fun hs _ hi => +theorem sSupIndep_iff_pairwiseDisjoint {s : Set α} : sSupIndep s ↔ s.PairwiseDisjoint id := + ⟨sSupIndep.pairwiseDisjoint, fun hs _ hi => disjoint_sSup_iff.2 fun _ hj => hs hi hj.1 <| Ne.symm hj.2⟩ -alias ⟨_, _root_.Set.PairwiseDisjoint.setIndependent⟩ := setIndependent_iff_pairwiseDisjoint +@[deprecated (since := "2024-11-24")] +alias setIndependent_iff_pairwiseDisjoint := sSupIndep_iff_pairwiseDisjoint -theorem independent_iff_pairwiseDisjoint {f : ι → α} : Independent f ↔ Pairwise (Disjoint on f) := - ⟨Independent.pairwiseDisjoint, fun hs _ => +alias ⟨_, _root_.Set.PairwiseDisjoint.sSupIndep⟩ := sSupIndep_iff_pairwiseDisjoint + +theorem iSupIndep_iff_pairwiseDisjoint {f : ι → α} : iSupIndep f ↔ Pairwise (Disjoint on f) := + ⟨iSupIndep.pairwiseDisjoint, fun hs _ => disjoint_iSup_iff.2 fun _ => disjoint_iSup_iff.2 fun hij => hs hij.symm⟩ -end CompleteLattice +@[deprecated (since := "2024-11-24")] +alias independent_iff_pairwiseDisjoint := iSupIndep_iff_pairwiseDisjoint end Frame diff --git a/Mathlib/RingTheory/FiniteLength.lean b/Mathlib/RingTheory/FiniteLength.lean index b1cf98eb550b8..8eea41b4e20e2 100644 --- a/Mathlib/RingTheory/FiniteLength.lean +++ b/Mathlib/RingTheory/FiniteLength.lean @@ -78,13 +78,13 @@ theorem isFiniteLength_iff_exists_compositionSeries : theorem IsSemisimpleModule.finite_tfae [IsSemisimpleModule R M] : List.TFAE [Module.Finite R M, IsNoetherian R M, IsArtinian R M, IsFiniteLength R M, - ∃ s : Set (Submodule R M), s.Finite ∧ CompleteLattice.SetIndependent s ∧ + ∃ s : Set (Submodule R M), s.Finite ∧ sSupIndep s ∧ sSup s = ⊤ ∧ ∀ m ∈ s, IsSimpleModule R m] := by rw [isFiniteLength_iff_isNoetherian_isArtinian] - obtain ⟨s, hs⟩ := IsSemisimpleModule.exists_setIndependent_sSup_simples_eq_top R M + obtain ⟨s, hs⟩ := IsSemisimpleModule.exists_sSupIndep_sSup_simples_eq_top R M tfae_have 1 ↔ 2 := ⟨fun _ ↦ inferInstance, fun _ ↦ inferInstance⟩ - tfae_have 2 → 5 := fun _ ↦ ⟨s, CompleteLattice.WellFoundedGT.finite_of_setIndependent hs.1, hs⟩ - tfae_have 3 → 5 := fun _ ↦ ⟨s, CompleteLattice.WellFoundedLT.finite_of_setIndependent hs.1, hs⟩ + tfae_have 2 → 5 := fun _ ↦ ⟨s, WellFoundedGT.finite_of_sSupIndep hs.1, hs⟩ + tfae_have 3 → 5 := fun _ ↦ ⟨s, WellFoundedLT.finite_of_sSupIndep hs.1, hs⟩ tfae_have 5 → 4 := fun ⟨s, fin, _, sSup_eq_top, simple⟩ ↦ by rw [← isNoetherian_top_iff, ← Submodule.topEquiv.isArtinian_iff, ← sSup_eq_top, sSup_eq_iSup, ← iSup_subtype''] diff --git a/Mathlib/RingTheory/Noetherian/Defs.lean b/Mathlib/RingTheory/Noetherian/Defs.lean index 953d4b7265411..926f9ba8c6880 100644 --- a/Mathlib/RingTheory/Noetherian/Defs.lean +++ b/Mathlib/RingTheory/Noetherian/Defs.lean @@ -313,17 +313,20 @@ universe w variable {R M P : Type*} {N : Type w} [Ring R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] [AddCommGroup P] [Module R P] [IsNoetherian R M] -lemma Submodule.finite_ne_bot_of_independent {ι : Type*} {N : ι → Submodule R M} - (h : CompleteLattice.Independent N) : +lemma Submodule.finite_ne_bot_of_iSupIndep {ι : Type*} {N : ι → Submodule R M} + (h : iSupIndep N) : Set.Finite {i | N i ≠ ⊥} := - CompleteLattice.WellFoundedGT.finite_ne_bot_of_independent h + WellFoundedGT.finite_ne_bot_of_iSupIndep h + +@[deprecated (since := "2024-11-24")] +alias Submodule.finite_ne_bot_of_independent := Submodule.finite_ne_bot_of_iSupIndep /-- A linearly-independent family of vectors in a module over a non-trivial ring must be finite if the module is Noetherian. -/ theorem LinearIndependent.finite_of_isNoetherian [Nontrivial R] {ι} {v : ι → M} (hv : LinearIndependent R v) : Finite ι := by - refine CompleteLattice.WellFoundedGT.finite_of_independent - hv.independent_span_singleton + refine WellFoundedGT.finite_of_iSupIndep + hv.iSupIndep_span_singleton fun i contra => ?_ apply hv.ne_zero i have : v i ∈ R ∙ v i := Submodule.mem_span_singleton_self (v i) diff --git a/Mathlib/RingTheory/SimpleModule.lean b/Mathlib/RingTheory/SimpleModule.lean index 66282ba413b64..dbae9bcb743ed 100644 --- a/Mathlib/RingTheory/SimpleModule.lean +++ b/Mathlib/RingTheory/SimpleModule.lean @@ -198,12 +198,14 @@ theorem exists_simple_submodule [Nontrivial M] : ∃ m : Submodule R M, IsSimple theorem sSup_simples_eq_top : sSup { m : Submodule R M | IsSimpleModule R m } = ⊤ := by simpa only [isSimpleModule_iff_isAtom] using sSup_atoms_eq_top -theorem exists_setIndependent_sSup_simples_eq_top : - ∃ s : Set (Submodule R M), CompleteLattice.SetIndependent s ∧ - sSup s = ⊤ ∧ ∀ m ∈ s, IsSimpleModule R m := by +theorem exists_sSupIndep_sSup_simples_eq_top : + ∃ s : Set (Submodule R M), sSupIndep s ∧ sSup s = ⊤ ∧ ∀ m ∈ s, IsSimpleModule R m := by have := sSup_simples_eq_top R M simp_rw [isSimpleModule_iff_isAtom] at this ⊢ - exact exists_setIndependent_of_sSup_atoms_eq_top this + exact exists_sSupIndep_of_sSup_atoms_eq_top this + +@[deprecated (since := "2024-11-24")] +alias exists_setIndependent_sSup_simples_eq_top := exists_sSupIndep_sSup_simples_eq_top /-- The annihilator of a semisimple module over a commutative ring is a radical ideal. -/ theorem annihilator_isRadical (R) [CommRing R] [Module R M] [IsSemisimpleModule R M] : diff --git a/scripts/nolints_prime_decls.txt b/scripts/nolints_prime_decls.txt index fd905374bfebb..a5ca83037daf6 100644 --- a/scripts/nolints_prime_decls.txt +++ b/scripts/nolints_prime_decls.txt @@ -686,11 +686,11 @@ CompactIccSpace.mk' CompactIccSpace.mk'' CompHaus.toProfinite_obj' compl_beattySeq' -CompleteLattice.Independent.comp' -CompleteLattice.independent_def' -CompleteLattice.independent_def'' -CompleteLattice.independent_of_dfinsupp_sumAddHom_injective' -CompleteLattice.Independent.supIndep' +iSupIndep.comp' +iSupIndep_def' +iSupIndep_def'' +iSupIndep_of_dfinsupp_sumAddHom_injective' +iSupIndep.supIndep' CompleteLattice.inf_continuous' CompleteLattice.sSup_continuous' CompletelyDistribLattice.MinimalAxioms.iInf_iSup_eq' @@ -2233,7 +2233,7 @@ LieIdeal.map_sup_ker_eq_map' LieModule.chainTop_isNonZero' LieModule.coe_chainTop' LieModule.genWeightSpaceChain_def' -LieModule.independent_genWeightSpace' +LieModule.iSupIndep_genWeightSpace' LieModule.instIsTrivialOfSubsingleton' LieModule.isNilpotent_of_top_iff' LieModule.iSup_genWeightSpace_eq_top' From b96544419e0739c6aefecfc4ee368400a77266ce Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 25 Nov 2024 10:13:40 +0000 Subject: [PATCH 276/829] doc: describe why map_ofNat can't be simp (yet) (#19410) See the benchmark in #19388, which while not disastrous is not ideal either. Unless the situation changes with https://github.com/leanprover/lean4/issues/2867, the DiscrTree key would apply to all funlike operations applied to any argument, not just applied to numerals. --- Mathlib/Data/Nat/Cast/Basic.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Mathlib/Data/Nat/Cast/Basic.lean b/Mathlib/Data/Nat/Cast/Basic.lean index 63025e62d2769..c494b68c05067 100644 --- a/Mathlib/Data/Nat/Cast/Basic.lean +++ b/Mathlib/Data/Nat/Cast/Basic.lean @@ -159,6 +159,11 @@ theorem eq_natCast [FunLike F ℕ R] [RingHomClass F ℕ R] (f : F) : ∀ n, f n theorem map_natCast [FunLike F R S] [RingHomClass F R S] (f : F) : ∀ n : ℕ, f (n : R) = n := map_natCast' f <| map_one f +/-- This lemma is not marked `@[simp]` lemma because its `#discr_tree_key` (for the LHS) would just +be `DFunLike.coe _ _`, due to the `no_index` that https://github.com/leanprover/lean4/issues/2867 +forces us to include, and therefore it would negatively impact performance. + +If that issue is resolved, this can be marked `@[simp]`. -/ theorem map_ofNat [FunLike F R S] [RingHomClass F R S] (f : F) (n : ℕ) [Nat.AtLeastTwo n] : (f (no_index (OfNat.ofNat n)) : S) = OfNat.ofNat n := map_natCast f n From a151e519ad7822530ddb3fb25275d41aa3a80c55 Mon Sep 17 00:00:00 2001 From: su00000 Date: Mon, 25 Nov 2024 11:04:38 +0000 Subject: [PATCH 277/829] feat(RingTheory/LocalProperties): Add theorems about LocalizedModule (#19080) This PR contains: 1. A proof that localization commutes with ranges. 2. Some missing lemmas about localized submodules. 3. A proof that injective, surjective, bijective (of linear maps) are local properties. These theorems will be used to prove that flat is a local property in a later PR. Co-authored-by: Yongle Hu Co-authored-by: Yi Song Co-authored-by: su00000 --- Mathlib.lean | 1 + .../Module/LocalizedModule/Submodule.lean | 56 ++--- Mathlib/RingTheory/Ideal/Maximal.lean | 9 + .../RingTheory/LocalProperties/Exactness.lean | 194 ++++++++++++++++++ .../RingTheory/LocalProperties/Submodule.lean | 81 +++++++- 5 files changed, 316 insertions(+), 25 deletions(-) create mode 100644 Mathlib/RingTheory/LocalProperties/Exactness.lean diff --git a/Mathlib.lean b/Mathlib.lean index 4b515136d2aa3..2c97f33123aff 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4322,6 +4322,7 @@ import Mathlib.RingTheory.LaurentSeries import Mathlib.RingTheory.LinearDisjoint import Mathlib.RingTheory.LittleWedderburn import Mathlib.RingTheory.LocalProperties.Basic +import Mathlib.RingTheory.LocalProperties.Exactness import Mathlib.RingTheory.LocalProperties.IntegrallyClosed import Mathlib.RingTheory.LocalProperties.Projective import Mathlib.RingTheory.LocalProperties.Reduced diff --git a/Mathlib/Algebra/Module/LocalizedModule/Submodule.lean b/Mathlib/Algebra/Module/LocalizedModule/Submodule.lean index 650844736f7ca..0ec7f25bf02d7 100644 --- a/Mathlib/Algebra/Module/LocalizedModule/Submodule.lean +++ b/Mathlib/Algebra/Module/LocalizedModule/Submodule.lean @@ -189,31 +189,32 @@ instance (M' : Submodule R M) : IsLocalizedModule p (M'.toLocalizedQuotient p) : end Quotient -section LinearMap +namespace LinearMap -variable {P : Type*} [AddCommGroup P] [Module R P] -variable {Q : Type*} [AddCommGroup Q] [Module R Q] [Module S Q] [IsScalarTower R S Q] +variable {P : Type*} [AddCommMonoid P] [Module R P] +variable {Q : Type*} [AddCommMonoid Q] [Module R Q] [Module S Q] [IsScalarTower R S Q] variable (f' : P →ₗ[R] Q) [IsLocalizedModule p f'] -lemma LinearMap.localized'_ker_eq_ker_localizedMap (g : M →ₗ[R] P) : - Submodule.localized' S p f (LinearMap.ker g) = - LinearMap.ker ((IsLocalizedModule.map p f f' g).extendScalarsOfIsLocalization p S) := by +open Submodule IsLocalizedModule + +lemma ker_localizedMap_eq_localized₀_ker (g : M →ₗ[R] P) : + ker (map p f f' g) = (ker g).localized₀ p f := by ext x - simp only [Submodule.mem_localized', mem_ker, extendScalarsOfIsLocalization_apply'] - constructor - · rintro ⟨m, hm, a, ha, rfl⟩ - rw [IsLocalizedModule.map_mk', hm] - simp - · intro h - obtain ⟨⟨a, b⟩, rfl⟩ := IsLocalizedModule.mk'_surjective p f x - simp only [Function.uncurry_apply_pair, IsLocalizedModule.map_mk', - IsLocalizedModule.mk'_eq_zero, IsLocalizedModule.eq_zero_iff p f'] at h + simp only [Submodule.mem_localized₀, mem_ker] + refine ⟨fun h ↦ ?_, ?_⟩ + · obtain ⟨⟨a, b⟩, rfl⟩ := IsLocalizedModule.mk'_surjective p f x + simp only [Function.uncurry_apply_pair, map_mk', mk'_eq_zero, eq_zero_iff p f'] at h obtain ⟨c, hc⟩ := h refine ⟨c • a, by simpa, c * b, by simp⟩ + · rintro ⟨m, hm, a, ha, rfl⟩ + simp [IsLocalizedModule.map_mk', hm] -lemma LinearMap.ker_localizedMap_eq_localized'_ker (g : M →ₗ[R] P) : - LinearMap.ker (IsLocalizedModule.map p f f' g) = - ((LinearMap.ker g).localized' S p f).restrictScalars _ := by +lemma localized'_ker_eq_ker_localizedMap (g : M →ₗ[R] P) : + (ker g).localized' S p f = ker ((map p f f' g).extendScalarsOfIsLocalization p S) := + SetLike.ext (by apply SetLike.ext_iff.mp (f.ker_localizedMap_eq_localized₀_ker p f' g).symm) + +lemma ker_localizedMap_eq_localized'_ker (g : M →ₗ[R] P) : + ker (map p f f' g) = ((ker g).localized' S p f).restrictScalars _ := by ext simp [localized'_ker_eq_ker_localizedMap S p f f'] @@ -223,18 +224,27 @@ The canonical map from the kernel of `g` to the kernel of `g` localized at a sub This is a localization map by `LinearMap.toKerLocalized_isLocalizedModule`. -/ @[simps!] -noncomputable def LinearMap.toKerIsLocalized (g : M →ₗ[R] P) : - ker g →ₗ[R] ker (IsLocalizedModule.map p f f' g) := - f.restrict (fun x hx ↦ by simp [LinearMap.mem_ker, LinearMap.mem_ker.mp hx]) +noncomputable def toKerIsLocalized (g : M →ₗ[R] P) : + ker g →ₗ[R] ker (map p f f' g) := + f.restrict (fun x hx ↦ by simp [mem_ker, mem_ker.mp hx]) include S in /-- The canonical map to the kernel of the localization of `g` is localizing. In other words, localization commutes with kernels. -/ -lemma LinearMap.toKerLocalized_isLocalizedModule (g : M →ₗ[R] P) : +lemma toKerLocalized_isLocalizedModule (g : M →ₗ[R] P) : IsLocalizedModule p (toKerIsLocalized p f f' g) := let e : Submodule.localized' S p f (ker g) ≃ₗ[S] - ker ((IsLocalizedModule.map p f f' g).extendScalarsOfIsLocalization p S) := + ker ((map p f f' g).extendScalarsOfIsLocalization p S) := LinearEquiv.ofEq _ _ (localized'_ker_eq_ker_localizedMap S p f f' g) IsLocalizedModule.of_linearEquiv p (Submodule.toLocalized' S p f (ker g)) (e.restrictScalars R) +lemma range_localizedMap_eq_localized₀_range (g : M →ₗ[R] P) : + range (map p f f' g) = (range g).localized₀ p f' := by + ext; simp [mem_localized₀, mem_range, (mk'_surjective p f).exists] + +/-- Localization commutes with ranges. -/ +lemma localized'_range_eq_range_localizedMap (g : M →ₗ[R] P) : + (range g).localized' S p f' = range ((map p f f' g).extendScalarsOfIsLocalization p S) := + SetLike.ext (by apply SetLike.ext_iff.mp (f.range_localizedMap_eq_localized₀_range p f' g).symm) + end LinearMap diff --git a/Mathlib/RingTheory/Ideal/Maximal.lean b/Mathlib/RingTheory/Ideal/Maximal.lean index 591bb5e34a064..444898cdd4f69 100644 --- a/Mathlib/RingTheory/Ideal/Maximal.lean +++ b/Mathlib/RingTheory/Ideal/Maximal.lean @@ -162,6 +162,15 @@ theorem IsMaximal.isPrime {I : Ideal α} (H : I.IsMaximal) : I.IsPrime := instance (priority := 100) IsMaximal.isPrime' (I : Ideal α) : ∀ [_H : I.IsMaximal], I.IsPrime := @IsMaximal.isPrime _ _ _ +theorem exists_disjoint_powers_of_span_eq_top (s : Set α) (hs : span s = ⊤) (I : Ideal α) + (hI : I ≠ ⊤) : ∃ r ∈ s, Disjoint (I : Set α) (Submonoid.powers r) := by + have ⟨M, hM, le⟩ := exists_le_maximal I hI + have := hM.1.1 + rw [Ne, eq_top_iff, ← hs, span_le, Set.not_subset] at this + have ⟨a, has, haM⟩ := this + exact ⟨a, has, Set.disjoint_left.mpr fun x hx ⟨n, hn⟩ ↦ + haM (hM.isPrime.mem_of_pow_mem _ (le <| hn ▸ hx))⟩ + theorem span_singleton_lt_span_singleton [IsDomain α] {x y : α} : span ({x} : Set α) < span ({y} : Set α) ↔ DvdNotUnit y x := by rw [lt_iff_le_not_le, span_singleton_le_span_singleton, span_singleton_le_span_singleton, diff --git a/Mathlib/RingTheory/LocalProperties/Exactness.lean b/Mathlib/RingTheory/LocalProperties/Exactness.lean new file mode 100644 index 0000000000000..3c6ae6bb837c7 --- /dev/null +++ b/Mathlib/RingTheory/LocalProperties/Exactness.lean @@ -0,0 +1,194 @@ +/- +Copyright (c) 2024 Sihan Su. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sihan Su, Yongle Hu, Yi Song +-/ +import Mathlib.Algebra.Exact +import Mathlib.RingTheory.LocalProperties.Submodule +import Mathlib.RingTheory.Localization.Away.Basic + +/-! +# Local properties about linear maps + +In this file, we show that +injectivity, surjectivity, bijectivity and exactness of linear maps are local properties. +More precisely, we show that these can be checked at maximal ideals and on standard covers. + +-/ +open Submodule LocalizedModule Ideal LinearMap + +section isLocalized_maximal + +open IsLocalizedModule + +variable {R M N L : Type*} [CommSemiring R] [AddCommMonoid M] [Module R M] + [AddCommMonoid N] [Module R N] [AddCommMonoid L] [Module R L] + +variable + (Mₚ : ∀ (P : Ideal R) [P.IsMaximal], Type*) + [∀ (P : Ideal R) [P.IsMaximal], AddCommMonoid (Mₚ P)] + [∀ (P : Ideal R) [P.IsMaximal], Module R (Mₚ P)] + (f : ∀ (P : Ideal R) [P.IsMaximal], M →ₗ[R] Mₚ P) + [∀ (P : Ideal R) [P.IsMaximal], IsLocalizedModule P.primeCompl (f P)] + (Nₚ : ∀ (P : Ideal R) [P.IsMaximal], Type*) + [∀ (P : Ideal R) [P.IsMaximal], AddCommMonoid (Nₚ P)] + [∀ (P : Ideal R) [P.IsMaximal], Module R (Nₚ P)] + (g : ∀ (P : Ideal R) [P.IsMaximal], N →ₗ[R] Nₚ P) + [∀ (P : Ideal R) [P.IsMaximal], IsLocalizedModule P.primeCompl (g P)] + (Lₚ : ∀ (P : Ideal R) [P.IsMaximal], Type*) + [∀ (P : Ideal R) [P.IsMaximal], AddCommMonoid (Lₚ P)] + [∀ (P : Ideal R) [P.IsMaximal], Module R (Lₚ P)] + (h : ∀ (P : Ideal R) [P.IsMaximal], L →ₗ[R] Lₚ P) + [∀ (P : Ideal R) [P.IsMaximal], IsLocalizedModule P.primeCompl (h P)] + (F : M →ₗ[R] N) (G : N →ₗ[R] L) + +theorem injective_of_isLocalized_maximal + (H : ∀ (P : Ideal R) [P.IsMaximal], Function.Injective (map P.primeCompl (f P) (g P) F)) : + Function.Injective F := + fun x y eq ↦ Module.eq_of_localization_maximal _ f _ _ fun P _ ↦ H P <| by simp [eq] + +theorem surjective_of_isLocalized_maximal + (H : ∀ (P : Ideal R) [P.IsMaximal], Function.Surjective (map P.primeCompl (f P) (g P) F)) : + Function.Surjective F := + range_eq_top.mp <| eq_top_of_localization₀_maximal Nₚ g _ <| + fun P _ ↦ (range_localizedMap_eq_localized₀_range _ (f P) (g P) F).symm.trans <| + range_eq_top.mpr <| H P + +theorem bijective_of_isLocalized_maximal + (H : ∀ (P : Ideal R) [P.IsMaximal], Function.Bijective (map P.primeCompl (f P) (g P) F)) : + Function.Bijective F := + ⟨injective_of_isLocalized_maximal Mₚ f Nₚ g F fun J _ ↦ (H J).1, + surjective_of_isLocalized_maximal Mₚ f Nₚ g F fun J _ ↦ (H J).2⟩ + +theorem exact_of_isLocalized_maximal (H : ∀ (J : Ideal R) [J.IsMaximal], + Function.Exact (map J.primeCompl (f J) (g J) F) (map J.primeCompl (g J) (h J) G)) : + Function.Exact F G := by + simp only [LinearMap.exact_iff] at H ⊢ + apply eq_of_localization₀_maximal Nₚ g + intro J hJ + rw [← LinearMap.range_localizedMap_eq_localized₀_range _ (f J) (g J) F, + ← LinearMap.ker_localizedMap_eq_localized₀_ker J.primeCompl (g J) (h J) G] + have := SetLike.ext_iff.mp <| H J + ext x + simp only [mem_range, mem_ker] at this ⊢ + exact this x + +end isLocalized_maximal + +section localized_maximal + +variable {R M N L : Type*} [CommSemiring R] [AddCommMonoid M] [Module R M] + [AddCommMonoid N] [Module R N] [AddCommMonoid L] [Module R L] (f : M →ₗ[R] N) (g : N →ₗ[R] L) + +theorem injective_of_localized_maximal + (h : ∀ (J : Ideal R) [J.IsMaximal], Function.Injective (map J.primeCompl f)) : + Function.Injective f := + injective_of_isLocalized_maximal _ (fun _ _ ↦ mkLinearMap _ _) _ (fun _ _ ↦ mkLinearMap _ _) f h + +theorem surjective_of_localized_maximal + (h : ∀ (J : Ideal R) [J.IsMaximal], Function.Surjective (map J.primeCompl f)) : + Function.Surjective f := + surjective_of_isLocalized_maximal _ (fun _ _ ↦ mkLinearMap _ _) _ (fun _ _ ↦ mkLinearMap _ _) f h + +theorem bijective_of_localized_maximal + (h : ∀ (J : Ideal R) [J.IsMaximal], Function.Bijective (map J.primeCompl f)) : + Function.Bijective f := + ⟨injective_of_localized_maximal _ fun J _ ↦ (h J).1, + surjective_of_localized_maximal _ fun J _ ↦ (h J).2⟩ + +theorem exact_of_localized_maximal + (h : ∀ (J : Ideal R) [J.IsMaximal], Function.Exact (map J.primeCompl f) (map J.primeCompl g)) : + Function.Exact f g := + exact_of_isLocalized_maximal _ (fun _ _ ↦ mkLinearMap _ _) _ (fun _ _ ↦ mkLinearMap _ _) + _ (fun _ _ ↦ mkLinearMap _ _) f g h + +end localized_maximal + +section isLocalized_span + +open IsLocalizedModule + +variable {R M N L : Type*} [CommSemiring R] [AddCommMonoid M] [Module R M] + [AddCommMonoid N] [Module R N] [AddCommMonoid L] [Module R L] (s : Set R) (spn : Ideal.span s = ⊤) +include spn + +variable + (Mₚ : ∀ _ : s, Type*) + [∀ r : s, AddCommMonoid (Mₚ r)] + [∀ r : s, Module R (Mₚ r)] + (f : ∀ r : s, M →ₗ[R] Mₚ r) + [∀ r : s, IsLocalizedModule (.powers r.1) (f r)] + (Nₚ : ∀ _ : s, Type*) + [∀ r : s, AddCommMonoid (Nₚ r)] + [∀ r : s, Module R (Nₚ r)] + (g : ∀ r : s, N →ₗ[R] Nₚ r) + [∀ r : s, IsLocalizedModule (.powers r.1) (g r)] + (Lₚ : ∀ _ : s, Type*) + [∀ r : s, AddCommMonoid (Lₚ r)] + [∀ r : s, Module R (Lₚ r)] + (h : ∀ r : s, L →ₗ[R] Lₚ r) + [∀ r : s, IsLocalizedModule (.powers r.1) (h r)] + (F : M →ₗ[R] N) (G : N →ₗ[R] L) + +theorem injective_of_isLocalized_span + (H : ∀ r : s, Function.Injective (map (.powers r.1) (f r) (g r) F)) : + Function.Injective F := + fun x y eq ↦ Module.eq_of_isLocalized_span _ spn _ f _ _ fun P ↦ H P <| by simp [eq] + +theorem surjective_of_isLocalized_span + (H : ∀ r : s, Function.Surjective (map (.powers r.1) (f r) (g r) F)) : + Function.Surjective F := + range_eq_top.mp <| eq_top_of_isLocalized₀_span s spn Nₚ g fun r ↦ + (range_localizedMap_eq_localized₀_range _ (f r) (g r) F).symm.trans <| range_eq_top.mpr <| H r + +theorem bijective_of_isLocalized_span + (H : ∀ r : s, Function.Bijective (map (.powers r.1) (f r) (g r) F)) : + Function.Bijective F := + ⟨injective_of_isLocalized_span _ spn Mₚ f Nₚ g F fun r ↦ (H r).1, + surjective_of_isLocalized_span _ spn Mₚ f Nₚ g F fun r ↦ (H r).2⟩ + +lemma exact_of_isLocalized_span (H : ∀ r : s, Function.Exact + (map (.powers r.1) (f r) (g r) F) (map (.powers r.1) (g r) (h r) G)) : + Function.Exact F G := by + simp only [LinearMap.exact_iff] at H ⊢ + apply Submodule.eq_of_isLocalized₀_span s spn Nₚ g + intro r + rw [← LinearMap.range_localizedMap_eq_localized₀_range _ (f r) (g r) F] + rw [← LinearMap.ker_localizedMap_eq_localized₀_ker (.powers r.1) (g r) (h r) G] + have := SetLike.ext_iff.mp <| H r + ext x + simp only [mem_range, mem_ker] at this ⊢ + exact this x + +end isLocalized_span + +section localized_span + +variable {R M N L : Type*} [CommSemiring R] [AddCommMonoid M] [Module R M] + [AddCommMonoid N] [Module R N] [AddCommMonoid L] [Module R L] + (s : Set R) (spn : span s = ⊤) (f : M →ₗ[R] N) (g : N →ₗ[R] L) +include spn + +theorem injective_of_localized_span + (h : ∀ r : s, Function.Injective (map (.powers r.1) f)) : + Function.Injective f := + injective_of_isLocalized_span s spn _ (fun _ ↦ mkLinearMap _ _) _ (fun _ ↦ mkLinearMap _ _) f h + +theorem surjective_of_localized_span + (h : ∀ r : s, Function.Surjective (map (.powers r.1) f)) : + Function.Surjective f := + surjective_of_isLocalized_span s spn _ (fun _ ↦ mkLinearMap _ _) _ (fun _ ↦ mkLinearMap _ _) f h + +theorem bijective_of_localized_span + (h : ∀ r : s, Function.Bijective (map (.powers r.1) f)) : + Function.Bijective f := + ⟨injective_of_localized_span _ spn _ fun r ↦ (h r).1, + surjective_of_localized_span _ spn _ fun r ↦ (h r).2⟩ + +lemma exact_of_localized_span + (h : ∀ r : s, Function.Exact (map (.powers r.1) f) (map (.powers r.1) g)) : + Function.Exact f g := + exact_of_isLocalized_span s spn _ (fun _ ↦ mkLinearMap _ _) _ (fun _ ↦ mkLinearMap _ _) + _ (fun _ ↦ mkLinearMap _ _) f g h + +end localized_span diff --git a/Mathlib/RingTheory/LocalProperties/Submodule.lean b/Mathlib/RingTheory/LocalProperties/Submodule.lean index 2ecaa20be887a..90ade41827edf 100644 --- a/Mathlib/RingTheory/LocalProperties/Submodule.lean +++ b/Mathlib/RingTheory/LocalProperties/Submodule.lean @@ -5,6 +5,7 @@ Authors: Andrew Yang, David Swinarski -/ import Mathlib.Algebra.Module.LocalizedModule.Submodule import Mathlib.RingTheory.Localization.AtPrime +import Mathlib.RingTheory.Localization.Away.Basic /-! # Local properties of modules and submodules @@ -80,8 +81,8 @@ theorem Submodule.eq_top_of_localization₀_maximal (N : Submodule R M) theorem Module.eq_of_localization_maximal (m m' : M) (h : ∀ (P : Ideal R) [P.IsMaximal], f P m = f P m') : m = m' := by - by_contra! ne - rw [← one_smul R m, ← one_smul R m'] at ne + rw [← one_smul R m, ← one_smul R m'] + by_contra ne have ⟨P, mP, le⟩ := (eqIdeal R m m').exists_le_maximal ((Ideal.ne_top_iff_one _).mpr ne) have ⟨s, hs⟩ := (IsLocalizedModule.eq_iff_exists P.primeCompl _).mp (h P) exact s.2 (le hs) @@ -124,3 +125,79 @@ theorem Submodule.eq_top_of_localization_maximal (N : Submodule R M) Submodule.eq_of_localization_maximal Rₚ Mₚ f fun P hP ↦ by simpa using h P end maximal + +section span + +open IsLocalizedModule LocalizedModule Ideal + +variable (s : Set R) (span_eq : Ideal.span s = ⊤) +include span_eq + +variable + (Rₚ : ∀ _ : s, Type*) + [∀ r : s, CommSemiring (Rₚ r)] + [∀ r : s, Algebra R (Rₚ r)] + [∀ r : s, IsLocalization.Away r.1 (Rₚ r)] + (Mₚ : ∀ _ : s, Type*) + [∀ r : s, AddCommMonoid (Mₚ r)] + [∀ r : s, Module R (Mₚ r)] + [∀ r : s, Module (Rₚ r) (Mₚ r)] + [∀ r : s, IsScalarTower R (Rₚ r) (Mₚ r)] + (f : ∀ r : s, M →ₗ[R] Mₚ r) + [∀ r : s, IsLocalizedModule (.powers r.1) (f r)] + +theorem Module.eq_of_isLocalized_span (x y : M) (h : ∀ r : s, f r x = f r y) : x = y := by + suffices Module.eqIdeal R x y = ⊤ by simpa [Module.eqIdeal] using (eq_top_iff_one _).mp this + by_contra ne + have ⟨r, hrs, disj⟩ := exists_disjoint_powers_of_span_eq_top s span_eq _ ne + let r : s := ⟨r, hrs⟩ + have ⟨⟨_, n, rfl⟩, eq⟩ := (IsLocalizedModule.eq_iff_exists (.powers r.1) _).mp (h r) + exact Set.disjoint_left.mp disj eq ⟨n, rfl⟩ + +theorem Module.eq_zero_of_isLocalized_span (x : M) (h : ∀ r : s, f r x = 0) : x = 0 := + eq_of_isLocalized_span s span_eq _ f x 0 <| by simpa only [map_zero] using h + +theorem Submodule.mem_of_isLocalized_span {m : M} {N : Submodule R M} + (h : ∀ r : s, f r m ∈ N.localized₀ (.powers r.1) (f r)) : m ∈ N := by + let I : Ideal R := N.comap (LinearMap.toSpanSingleton R M m) + suffices I = ⊤ by simpa [I] using I.eq_top_iff_one.mp this + by_contra! ne + have ⟨r, hrs, disj⟩ := exists_disjoint_powers_of_span_eq_top s span_eq _ ne + let r : s := ⟨r, hrs⟩ + obtain ⟨a, ha, t, e⟩ := h r + rw [← IsLocalizedModule.mk'_one (.powers r.1), IsLocalizedModule.mk'_eq_mk'_iff] at e + have ⟨u, hu⟩ := e + simp_rw [smul_smul] at hu + exact Set.disjoint_right.mp disj (u * t).2 (by apply hu ▸ smul_mem _ _ ha) + +theorem Submodule.le_of_isLocalized_span {N P : Submodule R M} + (h : ∀ r : s, N.localized₀ (.powers r.1) (f r) ≤ P.localized₀ (.powers r.1) (f r)) : N ≤ P := + fun m hm ↦ mem_of_isLocalized_span s span_eq _ f fun r ↦ h r ⟨m, hm, 1, by simp⟩ + +theorem Submodule.eq_of_isLocalized₀_span {N P : Submodule R M} + (h : ∀ r : s, N.localized₀ (.powers r.1) (f r) = P.localized₀ (.powers r.1) (f r)) : N = P := + le_antisymm (le_of_isLocalized_span s span_eq _ _ fun r ↦ (h r).le) + (le_of_isLocalized_span s span_eq _ _ fun r ↦ (h r).ge) + +theorem Submodule.eq_bot_of_isLocalized₀_span {N : Submodule R M} + (h : ∀ r : s, N.localized₀ (.powers r.1) (f r) = ⊥) : N = ⊥ := + eq_of_isLocalized₀_span s span_eq Mₚ f fun _ ↦ by simp only [h, Submodule.localized₀_bot] + +theorem Submodule.eq_top_of_isLocalized₀_span {N : Submodule R M} + (h : ∀ r : s, N.localized₀ (.powers r.1) (f r) = ⊤) : N = ⊤ := + eq_of_isLocalized₀_span s span_eq Mₚ f fun _ ↦ by simp only [h, Submodule.localized₀_top] + +theorem Submodule.eq_of_isLocalized'_span {N P : Submodule R M} + (h : ∀ r, N.localized' (Rₚ r) (.powers r.1) (f r) = P.localized' (Rₚ r) (.powers r.1) (f r)) : + N = P := + eq_of_isLocalized₀_span s span_eq _ f fun r ↦ congr(restrictScalars _ $(h r)) + +theorem Submodule.eq_bot_of_isLocalized'_span {N : Submodule R M} + (h : ∀ r : s, N.localized' (Rₚ r) (.powers r.1) (f r) = ⊥) : N = ⊥ := + eq_of_isLocalized'_span s span_eq Rₚ Mₚ f fun _ ↦ by simp only [h, Submodule.localized'_bot] + +theorem Submodule.eq_top_of_isLocalized'_span {N : Submodule R M} + (h : ∀ r : s, N.localized' (Rₚ r) (.powers r.1) (f r) = ⊤) : N = ⊤ := + eq_of_isLocalized'_span s span_eq Rₚ Mₚ f fun _ ↦ by simp only [h, Submodule.localized'_top] + +end span From a3580b7e38bd301d807a41f7be9be153566602bf Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Mon, 25 Nov 2024 11:04:40 +0000 Subject: [PATCH 278/829] chore(Order/ConditionallyCompleteLattice): split off `Defs.lean` from `Basic.lean` (#19277) --- Mathlib.lean | 1 + .../ConditionallyCompleteLattice/Basic.lean | 218 +-------------- .../ConditionallyCompleteLattice/Defs.lean | 251 ++++++++++++++++++ scripts/noshake.json | 2 + 4 files changed, 255 insertions(+), 217 deletions(-) create mode 100644 Mathlib/Order/ConditionallyCompleteLattice/Defs.lean diff --git a/Mathlib.lean b/Mathlib.lean index 2c97f33123aff..34090017c1344 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3918,6 +3918,7 @@ import Mathlib.Order.CompletePartialOrder import Mathlib.Order.CompleteSublattice import Mathlib.Order.Concept import Mathlib.Order.ConditionallyCompleteLattice.Basic +import Mathlib.Order.ConditionallyCompleteLattice.Defs import Mathlib.Order.ConditionallyCompleteLattice.Finset import Mathlib.Order.ConditionallyCompleteLattice.Group import Mathlib.Order.ConditionallyCompleteLattice.Indexed diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean index bf1aa9e9ffaec..437a1e3d52e13 100644 --- a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean +++ b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean @@ -3,11 +3,8 @@ Copyright (c) 2018 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import Mathlib.Order.Bounds.Basic -import Mathlib.Order.WellFounded -import Mathlib.Data.Set.Image -import Mathlib.Order.Interval.Set.Basic import Mathlib.Data.Set.Lattice +import Mathlib.Order.ConditionallyCompleteLattice.Defs /-! # Theory of conditionally complete lattices @@ -119,50 +116,6 @@ theorem WithBot.coe_sInf' [InfSet α] {s : Set α} (hs : BddBelow s) : end -/-- A conditionally complete lattice is a lattice in which -every nonempty subset which is bounded above has a supremum, and -every nonempty subset which is bounded below has an infimum. -Typical examples are real numbers or natural numbers. - -To differentiate the statements from the corresponding statements in (unconditional) -complete lattices, we prefix `sInf` and `sSup` by a `c` everywhere. The same statements should -hold in both worlds, sometimes with additional assumptions of nonemptiness or -boundedness. -/ -class ConditionallyCompleteLattice (α : Type*) extends Lattice α, SupSet α, InfSet α where - /-- `a ≤ sSup s` for all `a ∈ s`. -/ - le_csSup : ∀ s a, BddAbove s → a ∈ s → a ≤ sSup s - /-- `sSup s ≤ a` for all `a ∈ upperBounds s`. -/ - csSup_le : ∀ s a, Set.Nonempty s → a ∈ upperBounds s → sSup s ≤ a - /-- `sInf s ≤ a` for all `a ∈ s`. -/ - csInf_le : ∀ s a, BddBelow s → a ∈ s → sInf s ≤ a - /-- `a ≤ sInf s` for all `a ∈ lowerBounds s`. -/ - le_csInf : ∀ s a, Set.Nonempty s → a ∈ lowerBounds s → a ≤ sInf s - --- Porting note: mathlib3 used `renaming` -/-- A conditionally complete linear order is a linear order in which -every nonempty subset which is bounded above has a supremum, and -every nonempty subset which is bounded below has an infimum. -Typical examples are real numbers or natural numbers. - -To differentiate the statements from the corresponding statements in (unconditional) -complete linear orders, we prefix `sInf` and `sSup` by a `c` everywhere. The same statements should -hold in both worlds, sometimes with additional assumptions of nonemptiness or -boundedness. -/ -class ConditionallyCompleteLinearOrder (α : Type*) extends ConditionallyCompleteLattice α where - /-- A `ConditionallyCompleteLinearOrder` is total. -/ - le_total (a b : α) : a ≤ b ∨ b ≤ a - /-- In a `ConditionallyCompleteLinearOrder`, we assume the order relations are all decidable. -/ - decidableLE : DecidableRel (· ≤ · : α → α → Prop) - /-- In a `ConditionallyCompleteLinearOrder`, we assume the order relations are all decidable. -/ - decidableEq : DecidableEq α := @decidableEqOfDecidableLE _ _ decidableLE - /-- In a `ConditionallyCompleteLinearOrder`, we assume the order relations are all decidable. -/ - decidableLT : DecidableRel (· < · : α → α → Prop) := - @decidableLTOfDecidableLE _ _ decidableLE - /-- If a set is not bounded above, its supremum is by convention `sSup ∅`. -/ - csSup_of_not_bddAbove : ∀ s, ¬BddAbove s → sSup s = sSup (∅ : Set α) - /-- If a set is not bounded below, its infimum is by convention `sInf ∅`. -/ - csInf_of_not_bddBelow : ∀ s, ¬BddBelow s → sInf s = sInf (∅ : Set α) - instance ConditionallyCompleteLinearOrder.toLinearOrder [ConditionallyCompleteLinearOrder α] : LinearOrder α := { ‹ConditionallyCompleteLinearOrder α› with @@ -179,19 +132,6 @@ instance ConditionallyCompleteLinearOrder.toLinearOrder [ConditionallyCompleteLi · simp [h₁] · simp [show ¬(a ≤ b) from fun h => hab (le_antisymm h h₂), h₂] } -/-- A conditionally complete linear order with `Bot` is a linear order with least element, in which -every nonempty subset which is bounded above has a supremum, and every nonempty subset (necessarily -bounded below) has an infimum. A typical example is the natural numbers. - -To differentiate the statements from the corresponding statements in (unconditional) -complete linear orders, we prefix `sInf` and `sSup` by a `c` everywhere. The same statements should -hold in both worlds, sometimes with additional assumptions of nonemptiness or -boundedness. -/ -class ConditionallyCompleteLinearOrderBot (α : Type*) extends ConditionallyCompleteLinearOrder α, - OrderBot α where - /-- The supremum of the empty set is special-cased to `⊥` -/ - csSup_empty : sSup ∅ = ⊥ - -- see Note [lower instance priority] attribute [instance 100] ConditionallyCompleteLinearOrderBot.toOrderBot @@ -214,36 +154,6 @@ instance (priority := 100) CompleteLinearOrder.toConditionallyCompleteLinearOrde csSup_of_not_bddAbove := fun s H ↦ (H (OrderTop.bddAbove s)).elim csInf_of_not_bddBelow := fun s H ↦ (H (OrderBot.bddBelow s)).elim } -open scoped Classical in -/-- A well founded linear order is conditionally complete, with a bottom element. -/ -noncomputable abbrev WellFoundedLT.conditionallyCompleteLinearOrderBot (α : Type*) - [i₁ : LinearOrder α] [i₂ : OrderBot α] [h : WellFoundedLT α] : - ConditionallyCompleteLinearOrderBot α := - { i₁, i₂, LinearOrder.toLattice with - sInf := fun s => if hs : s.Nonempty then h.wf.min s hs else ⊥ - csInf_le := fun s a _ has => by - have s_ne : s.Nonempty := ⟨a, has⟩ - simpa [s_ne] using not_lt.1 (h.wf.not_lt_min s s_ne has) - le_csInf := fun s a hs has => by - simp only [hs, dif_pos] - exact has (h.wf.min_mem s hs) - sSup := fun s => if hs : (upperBounds s).Nonempty then h.wf.min _ hs else ⊥ - le_csSup := fun s a hs has => by - have h's : (upperBounds s).Nonempty := hs - simp only [h's, dif_pos] - exact h.wf.min_mem _ h's has - csSup_le := fun s a _ has => by - have h's : (upperBounds s).Nonempty := ⟨a, has⟩ - simp only [h's, dif_pos] - simpa using h.wf.not_lt_min _ h's has - csSup_empty := by simpa using eq_bot_iff.2 (not_lt.1 <| h.wf.not_lt_min _ _ <| mem_univ ⊥) - csSup_of_not_bddAbove := by - intro s H - have B : ¬((upperBounds s).Nonempty) := H - simp only [B, dite_false, upperBounds_empty, univ_nonempty, dite_true] - exact le_antisymm bot_le (WellFounded.min_le _ (mem_univ _)) - csInf_of_not_bddBelow := fun s H ↦ (H (OrderBot.bddBelow s)).elim } - namespace OrderDual instance instConditionallyCompleteLattice (α : Type*) [ConditionallyCompleteLattice α] : @@ -261,132 +171,6 @@ instance (α : Type*) [ConditionallyCompleteLinearOrder α] : ConditionallyCompl end OrderDual -/-- Create a `ConditionallyCompleteLattice` from a `PartialOrder` and `sup` function -that returns the least upper bound of a nonempty set which is bounded above. Usually this -constructor provides poor definitional equalities. If other fields are known explicitly, they -should be provided; for example, if `inf` is known explicitly, construct the -`ConditionallyCompleteLattice` instance as -``` -instance : ConditionallyCompleteLattice my_T := - { inf := better_inf, - le_inf := ..., - inf_le_right := ..., - inf_le_left := ... - -- don't care to fix sup, sInf - ..conditionallyCompleteLatticeOfsSup my_T _ } -``` --/ -def conditionallyCompleteLatticeOfsSup (α : Type*) [H1 : PartialOrder α] [H2 : SupSet α] - (bddAbove_pair : ∀ a b : α, BddAbove ({a, b} : Set α)) - (bddBelow_pair : ∀ a b : α, BddBelow ({a, b} : Set α)) - (isLUB_sSup : ∀ s : Set α, BddAbove s → s.Nonempty → IsLUB s (sSup s)) : - ConditionallyCompleteLattice α := - { H1, H2 with - sup := fun a b => sSup {a, b} - le_sup_left := fun a b => - (isLUB_sSup {a, b} (bddAbove_pair a b) (insert_nonempty _ _)).1 (mem_insert _ _) - le_sup_right := fun a b => - (isLUB_sSup {a, b} (bddAbove_pair a b) (insert_nonempty _ _)).1 - (mem_insert_of_mem _ (mem_singleton _)) - sup_le := fun a b _ hac hbc => - (isLUB_sSup {a, b} (bddAbove_pair a b) (insert_nonempty _ _)).2 - (forall_insert_of_forall (forall_eq.mpr hbc) hac) - inf := fun a b => sSup (lowerBounds {a, b}) - inf_le_left := fun a b => - (isLUB_sSup (lowerBounds {a, b}) (Nonempty.bddAbove_lowerBounds ⟨a, mem_insert _ _⟩) - (bddBelow_pair a b)).2 - fun _ hc => hc <| mem_insert _ _ - inf_le_right := fun a b => - (isLUB_sSup (lowerBounds {a, b}) (Nonempty.bddAbove_lowerBounds ⟨a, mem_insert _ _⟩) - (bddBelow_pair a b)).2 - fun _ hc => hc <| mem_insert_of_mem _ (mem_singleton _) - le_inf := fun c a b hca hcb => - (isLUB_sSup (lowerBounds {a, b}) (Nonempty.bddAbove_lowerBounds ⟨a, mem_insert _ _⟩) - ⟨c, forall_insert_of_forall (forall_eq.mpr hcb) hca⟩).1 - (forall_insert_of_forall (forall_eq.mpr hcb) hca) - sInf := fun s => sSup (lowerBounds s) - csSup_le := fun s a hs ha => (isLUB_sSup s ⟨a, ha⟩ hs).2 ha - le_csSup := fun s a hs ha => (isLUB_sSup s hs ⟨a, ha⟩).1 ha - csInf_le := fun s a hs ha => - (isLUB_sSup (lowerBounds s) (Nonempty.bddAbove_lowerBounds ⟨a, ha⟩) hs).2 fun _ hb => hb ha - le_csInf := fun s a hs ha => - (isLUB_sSup (lowerBounds s) hs.bddAbove_lowerBounds ⟨a, ha⟩).1 ha } - -/-- Create a `ConditionallyCompleteLattice` from a `PartialOrder` and `inf` function -that returns the greatest lower bound of a nonempty set which is bounded below. Usually this -constructor provides poor definitional equalities. If other fields are known explicitly, they -should be provided; for example, if `inf` is known explicitly, construct the -`ConditionallyCompleteLattice` instance as -``` -instance : ConditionallyCompleteLattice my_T := - { inf := better_inf, - le_inf := ..., - inf_le_right := ..., - inf_le_left := ... - -- don't care to fix sup, sSup - ..conditionallyCompleteLatticeOfsInf my_T _ } -``` --/ -def conditionallyCompleteLatticeOfsInf (α : Type*) [H1 : PartialOrder α] [H2 : InfSet α] - (bddAbove_pair : ∀ a b : α, BddAbove ({a, b} : Set α)) - (bddBelow_pair : ∀ a b : α, BddBelow ({a, b} : Set α)) - (isGLB_sInf : ∀ s : Set α, BddBelow s → s.Nonempty → IsGLB s (sInf s)) : - ConditionallyCompleteLattice α := - { H1, H2 with - inf := fun a b => sInf {a, b} - inf_le_left := fun a b => - (isGLB_sInf {a, b} (bddBelow_pair a b) (insert_nonempty _ _)).1 (mem_insert _ _) - inf_le_right := fun a b => - (isGLB_sInf {a, b} (bddBelow_pair a b) (insert_nonempty _ _)).1 - (mem_insert_of_mem _ (mem_singleton _)) - le_inf := fun _ a b hca hcb => - (isGLB_sInf {a, b} (bddBelow_pair a b) (insert_nonempty _ _)).2 - (forall_insert_of_forall (forall_eq.mpr hcb) hca) - sup := fun a b => sInf (upperBounds {a, b}) - le_sup_left := fun a b => - (isGLB_sInf (upperBounds {a, b}) (Nonempty.bddBelow_upperBounds ⟨a, mem_insert _ _⟩) - (bddAbove_pair a b)).2 - fun _ hc => hc <| mem_insert _ _ - le_sup_right := fun a b => - (isGLB_sInf (upperBounds {a, b}) (Nonempty.bddBelow_upperBounds ⟨a, mem_insert _ _⟩) - (bddAbove_pair a b)).2 - fun _ hc => hc <| mem_insert_of_mem _ (mem_singleton _) - sup_le := fun a b c hac hbc => - (isGLB_sInf (upperBounds {a, b}) (Nonempty.bddBelow_upperBounds ⟨a, mem_insert _ _⟩) - ⟨c, forall_insert_of_forall (forall_eq.mpr hbc) hac⟩).1 - (forall_insert_of_forall (forall_eq.mpr hbc) hac) - sSup := fun s => sInf (upperBounds s) - le_csInf := fun s a hs ha => (isGLB_sInf s ⟨a, ha⟩ hs).2 ha - csInf_le := fun s a hs ha => (isGLB_sInf s hs ⟨a, ha⟩).1 ha - le_csSup := fun s a hs ha => - (isGLB_sInf (upperBounds s) (Nonempty.bddBelow_upperBounds ⟨a, ha⟩) hs).2 fun _ hb => hb ha - csSup_le := fun s a hs ha => - (isGLB_sInf (upperBounds s) hs.bddBelow_upperBounds ⟨a, ha⟩).1 ha } - -/-- A version of `conditionallyCompleteLatticeOfsSup` when we already know that `α` is a lattice. - -This should only be used when it is both hard and unnecessary to provide `inf` explicitly. -/ -def conditionallyCompleteLatticeOfLatticeOfsSup (α : Type*) [H1 : Lattice α] [SupSet α] - (isLUB_sSup : ∀ s : Set α, BddAbove s → s.Nonempty → IsLUB s (sSup s)) : - ConditionallyCompleteLattice α := - { H1, - conditionallyCompleteLatticeOfsSup α - (fun a b => ⟨a ⊔ b, forall_insert_of_forall (forall_eq.mpr le_sup_right) le_sup_left⟩) - (fun a b => ⟨a ⊓ b, forall_insert_of_forall (forall_eq.mpr inf_le_right) inf_le_left⟩) - isLUB_sSup with } - -/-- A version of `conditionallyCompleteLatticeOfsInf` when we already know that `α` is a lattice. - -This should only be used when it is both hard and unnecessary to provide `sup` explicitly. -/ -def conditionallyCompleteLatticeOfLatticeOfsInf (α : Type*) [H1 : Lattice α] [InfSet α] - (isGLB_sInf : ∀ s : Set α, BddBelow s → s.Nonempty → IsGLB s (sInf s)) : - ConditionallyCompleteLattice α := - { H1, - conditionallyCompleteLatticeOfsInf α - (fun a b => ⟨a ⊔ b, forall_insert_of_forall (forall_eq.mpr le_sup_right) le_sup_left⟩) - (fun a b => ⟨a ⊓ b, forall_insert_of_forall (forall_eq.mpr inf_le_right) inf_le_left⟩) - isGLB_sInf with } - section ConditionallyCompleteLattice variable [ConditionallyCompleteLattice α] {s t : Set α} {a b : α} diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Defs.lean b/Mathlib/Order/ConditionallyCompleteLattice/Defs.lean new file mode 100644 index 0000000000000..4a499bbdfdc7a --- /dev/null +++ b/Mathlib/Order/ConditionallyCompleteLattice/Defs.lean @@ -0,0 +1,251 @@ +/- +Copyright (c) 2018 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel +-/ +import Mathlib.Order.Bounds.Basic +import Mathlib.Order.SetNotation +import Mathlib.Order.WellFounded + +/-! +# Definitions of conditionally complete lattices + +A conditionally complete lattice is a lattice in which every non-empty bounded subset `s` +has a least upper bound and a greatest lower bound, denoted below by `sSup s` and `sInf s`. +Typical examples are `ℝ`, `ℕ`, and `ℤ` with their usual orders. + +The theory is very comparable to the theory of complete lattices, except that suitable +boundedness and nonemptiness assumptions have to be added to most statements. +We express these using the `BddAbove` and `BddBelow` predicates, which we use to prove +most useful properties of `sSup` and `sInf` in conditionally complete lattices. + +To differentiate the statements between complete lattices and conditionally complete +lattices, we prefix `sInf` and `sSup` in the statements by `c`, giving `csInf` and `csSup`. +For instance, `sInf_le` is a statement in complete lattices ensuring `sInf s ≤ x`, +while `csInf_le` is the same statement in conditionally complete lattices +with an additional assumption that `s` is bounded below. +-/ + +open Set + +variable {α β γ : Type*} {ι : Sort*} + +/-- A conditionally complete lattice is a lattice in which +every nonempty subset which is bounded above has a supremum, and +every nonempty subset which is bounded below has an infimum. +Typical examples are real numbers or natural numbers. + +To differentiate the statements from the corresponding statements in (unconditional) +complete lattices, we prefix `sInf` and `sSup` by a `c` everywhere. The same statements should +hold in both worlds, sometimes with additional assumptions of nonemptiness or +boundedness. -/ +class ConditionallyCompleteLattice (α : Type*) extends Lattice α, SupSet α, InfSet α where + /-- `a ≤ sSup s` for all `a ∈ s`. -/ + le_csSup : ∀ s a, BddAbove s → a ∈ s → a ≤ sSup s + /-- `sSup s ≤ a` for all `a ∈ upperBounds s`. -/ + csSup_le : ∀ s a, Set.Nonempty s → a ∈ upperBounds s → sSup s ≤ a + /-- `sInf s ≤ a` for all `a ∈ s`. -/ + csInf_le : ∀ s a, BddBelow s → a ∈ s → sInf s ≤ a + /-- `a ≤ sInf s` for all `a ∈ lowerBounds s`. -/ + le_csInf : ∀ s a, Set.Nonempty s → a ∈ lowerBounds s → a ≤ sInf s + +-- Porting note: mathlib3 used `renaming` +/-- A conditionally complete linear order is a linear order in which +every nonempty subset which is bounded above has a supremum, and +every nonempty subset which is bounded below has an infimum. +Typical examples are real numbers or natural numbers. + +To differentiate the statements from the corresponding statements in (unconditional) +complete linear orders, we prefix `sInf` and `sSup` by a `c` everywhere. The same statements should +hold in both worlds, sometimes with additional assumptions of nonemptiness or +boundedness. -/ +class ConditionallyCompleteLinearOrder (α : Type*) extends ConditionallyCompleteLattice α where + /-- A `ConditionallyCompleteLinearOrder` is total. -/ + le_total (a b : α) : a ≤ b ∨ b ≤ a + /-- In a `ConditionallyCompleteLinearOrder`, we assume the order relations are all decidable. -/ + decidableLE : DecidableRel (· ≤ · : α → α → Prop) + /-- In a `ConditionallyCompleteLinearOrder`, we assume the order relations are all decidable. -/ + decidableEq : DecidableEq α := @decidableEqOfDecidableLE _ _ decidableLE + /-- In a `ConditionallyCompleteLinearOrder`, we assume the order relations are all decidable. -/ + decidableLT : DecidableRel (· < · : α → α → Prop) := + @decidableLTOfDecidableLE _ _ decidableLE + /-- If a set is not bounded above, its supremum is by convention `sSup ∅`. -/ + csSup_of_not_bddAbove : ∀ s, ¬BddAbove s → sSup s = sSup (∅ : Set α) + /-- If a set is not bounded below, its infimum is by convention `sInf ∅`. -/ + csInf_of_not_bddBelow : ∀ s, ¬BddBelow s → sInf s = sInf (∅ : Set α) + +/-- A conditionally complete linear order with `Bot` is a linear order with least element, in which +every nonempty subset which is bounded above has a supremum, and every nonempty subset (necessarily +bounded below) has an infimum. A typical example is the natural numbers. + +To differentiate the statements from the corresponding statements in (unconditional) +complete linear orders, we prefix `sInf` and `sSup` by a `c` everywhere. The same statements should +hold in both worlds, sometimes with additional assumptions of nonemptiness or +boundedness. -/ +class ConditionallyCompleteLinearOrderBot (α : Type*) extends ConditionallyCompleteLinearOrder α, + OrderBot α where + /-- The supremum of the empty set is special-cased to `⊥` -/ + csSup_empty : sSup ∅ = ⊥ + +-- see Note [lower instance priority] +attribute [instance 100] ConditionallyCompleteLinearOrderBot.toOrderBot + +open scoped Classical in +/-- A well founded linear order is conditionally complete, with a bottom element. -/ +noncomputable abbrev WellFoundedLT.conditionallyCompleteLinearOrderBot (α : Type*) + [i₁ : LinearOrder α] [i₂ : OrderBot α] [h : WellFoundedLT α] : + ConditionallyCompleteLinearOrderBot α := + { i₁, i₂, LinearOrder.toLattice with + sInf := fun s => if hs : s.Nonempty then h.wf.min s hs else ⊥ + csInf_le := fun s a _ has => by + have s_ne : s.Nonempty := ⟨a, has⟩ + simpa [s_ne] using not_lt.1 (h.wf.not_lt_min s s_ne has) + le_csInf := fun s a hs has => by + simp only [hs, dif_pos] + exact has (h.wf.min_mem s hs) + sSup := fun s => if hs : (upperBounds s).Nonempty then h.wf.min _ hs else ⊥ + le_csSup := fun s a hs has => by + have h's : (upperBounds s).Nonempty := hs + simp only [h's, dif_pos] + exact h.wf.min_mem _ h's has + csSup_le := fun s a _ has => by + have h's : (upperBounds s).Nonempty := ⟨a, has⟩ + simp only [h's, dif_pos] + simpa using h.wf.not_lt_min _ h's has + csSup_empty := by simpa using eq_bot_iff.2 (not_lt.1 <| h.wf.not_lt_min _ _ <| mem_univ ⊥) + csSup_of_not_bddAbove := by + intro s H + have B : ¬((upperBounds s).Nonempty) := H + simp only [B, dite_false, upperBounds_empty, univ_nonempty, dite_true] + exact le_antisymm bot_le (WellFounded.min_le _ (mem_univ _)) + csInf_of_not_bddBelow := fun s H ↦ (H (OrderBot.bddBelow s)).elim } + +namespace OrderDual + +end OrderDual + +/-- Create a `ConditionallyCompleteLattice` from a `PartialOrder` and `sup` function +that returns the least upper bound of a nonempty set which is bounded above. Usually this +constructor provides poor definitional equalities. If other fields are known explicitly, they +should be provided; for example, if `inf` is known explicitly, construct the +`ConditionallyCompleteLattice` instance as +``` +instance : ConditionallyCompleteLattice my_T := + { inf := better_inf, + le_inf := ..., + inf_le_right := ..., + inf_le_left := ... + -- don't care to fix sup, sInf + ..conditionallyCompleteLatticeOfsSup my_T _ } +``` +-/ +def conditionallyCompleteLatticeOfsSup (α : Type*) [H1 : PartialOrder α] [H2 : SupSet α] + (bddAbove_pair : ∀ a b : α, BddAbove ({a, b} : Set α)) + (bddBelow_pair : ∀ a b : α, BddBelow ({a, b} : Set α)) + (isLUB_sSup : ∀ s : Set α, BddAbove s → s.Nonempty → IsLUB s (sSup s)) : + ConditionallyCompleteLattice α := + { H1, H2 with + sup := fun a b => sSup {a, b} + le_sup_left := fun a b => + (isLUB_sSup {a, b} (bddAbove_pair a b) (insert_nonempty _ _)).1 (mem_insert _ _) + le_sup_right := fun a b => + (isLUB_sSup {a, b} (bddAbove_pair a b) (insert_nonempty _ _)).1 + (mem_insert_of_mem _ (mem_singleton _)) + sup_le := fun a b _ hac hbc => + (isLUB_sSup {a, b} (bddAbove_pair a b) (insert_nonempty _ _)).2 + (forall_insert_of_forall (forall_eq.mpr hbc) hac) + inf := fun a b => sSup (lowerBounds {a, b}) + inf_le_left := fun a b => + (isLUB_sSup (lowerBounds {a, b}) (Nonempty.bddAbove_lowerBounds ⟨a, mem_insert _ _⟩) + (bddBelow_pair a b)).2 + fun _ hc => hc <| mem_insert _ _ + inf_le_right := fun a b => + (isLUB_sSup (lowerBounds {a, b}) (Nonempty.bddAbove_lowerBounds ⟨a, mem_insert _ _⟩) + (bddBelow_pair a b)).2 + fun _ hc => hc <| mem_insert_of_mem _ (mem_singleton _) + le_inf := fun c a b hca hcb => + (isLUB_sSup (lowerBounds {a, b}) (Nonempty.bddAbove_lowerBounds ⟨a, mem_insert _ _⟩) + ⟨c, forall_insert_of_forall (forall_eq.mpr hcb) hca⟩).1 + (forall_insert_of_forall (forall_eq.mpr hcb) hca) + sInf := fun s => sSup (lowerBounds s) + csSup_le := fun s a hs ha => (isLUB_sSup s ⟨a, ha⟩ hs).2 ha + le_csSup := fun s a hs ha => (isLUB_sSup s hs ⟨a, ha⟩).1 ha + csInf_le := fun s a hs ha => + (isLUB_sSup (lowerBounds s) (Nonempty.bddAbove_lowerBounds ⟨a, ha⟩) hs).2 fun _ hb => hb ha + le_csInf := fun s a hs ha => + (isLUB_sSup (lowerBounds s) hs.bddAbove_lowerBounds ⟨a, ha⟩).1 ha } + +/-- Create a `ConditionallyCompleteLattice` from a `PartialOrder` and `inf` function +that returns the greatest lower bound of a nonempty set which is bounded below. Usually this +constructor provides poor definitional equalities. If other fields are known explicitly, they +should be provided; for example, if `inf` is known explicitly, construct the +`ConditionallyCompleteLattice` instance as +``` +instance : ConditionallyCompleteLattice my_T := + { inf := better_inf, + le_inf := ..., + inf_le_right := ..., + inf_le_left := ... + -- don't care to fix sup, sSup + ..conditionallyCompleteLatticeOfsInf my_T _ } +``` +-/ +def conditionallyCompleteLatticeOfsInf (α : Type*) [H1 : PartialOrder α] [H2 : InfSet α] + (bddAbove_pair : ∀ a b : α, BddAbove ({a, b} : Set α)) + (bddBelow_pair : ∀ a b : α, BddBelow ({a, b} : Set α)) + (isGLB_sInf : ∀ s : Set α, BddBelow s → s.Nonempty → IsGLB s (sInf s)) : + ConditionallyCompleteLattice α := + { H1, H2 with + inf := fun a b => sInf {a, b} + inf_le_left := fun a b => + (isGLB_sInf {a, b} (bddBelow_pair a b) (insert_nonempty _ _)).1 (mem_insert _ _) + inf_le_right := fun a b => + (isGLB_sInf {a, b} (bddBelow_pair a b) (insert_nonempty _ _)).1 + (mem_insert_of_mem _ (mem_singleton _)) + le_inf := fun _ a b hca hcb => + (isGLB_sInf {a, b} (bddBelow_pair a b) (insert_nonempty _ _)).2 + (forall_insert_of_forall (forall_eq.mpr hcb) hca) + sup := fun a b => sInf (upperBounds {a, b}) + le_sup_left := fun a b => + (isGLB_sInf (upperBounds {a, b}) (Nonempty.bddBelow_upperBounds ⟨a, mem_insert _ _⟩) + (bddAbove_pair a b)).2 + fun _ hc => hc <| mem_insert _ _ + le_sup_right := fun a b => + (isGLB_sInf (upperBounds {a, b}) (Nonempty.bddBelow_upperBounds ⟨a, mem_insert _ _⟩) + (bddAbove_pair a b)).2 + fun _ hc => hc <| mem_insert_of_mem _ (mem_singleton _) + sup_le := fun a b c hac hbc => + (isGLB_sInf (upperBounds {a, b}) (Nonempty.bddBelow_upperBounds ⟨a, mem_insert _ _⟩) + ⟨c, forall_insert_of_forall (forall_eq.mpr hbc) hac⟩).1 + (forall_insert_of_forall (forall_eq.mpr hbc) hac) + sSup := fun s => sInf (upperBounds s) + le_csInf := fun s a hs ha => (isGLB_sInf s ⟨a, ha⟩ hs).2 ha + csInf_le := fun s a hs ha => (isGLB_sInf s hs ⟨a, ha⟩).1 ha + le_csSup := fun s a hs ha => + (isGLB_sInf (upperBounds s) (Nonempty.bddBelow_upperBounds ⟨a, ha⟩) hs).2 fun _ hb => hb ha + csSup_le := fun s a hs ha => + (isGLB_sInf (upperBounds s) hs.bddBelow_upperBounds ⟨a, ha⟩).1 ha } + +/-- A version of `conditionallyCompleteLatticeOfsSup` when we already know that `α` is a lattice. + +This should only be used when it is both hard and unnecessary to provide `inf` explicitly. -/ +def conditionallyCompleteLatticeOfLatticeOfsSup (α : Type*) [H1 : Lattice α] [SupSet α] + (isLUB_sSup : ∀ s : Set α, BddAbove s → s.Nonempty → IsLUB s (sSup s)) : + ConditionallyCompleteLattice α := + { H1, + conditionallyCompleteLatticeOfsSup α + (fun a b => ⟨a ⊔ b, forall_insert_of_forall (forall_eq.mpr le_sup_right) le_sup_left⟩) + (fun a b => ⟨a ⊓ b, forall_insert_of_forall (forall_eq.mpr inf_le_right) inf_le_left⟩) + isLUB_sSup with } + +/-- A version of `conditionallyCompleteLatticeOfsInf` when we already know that `α` is a lattice. + +This should only be used when it is both hard and unnecessary to provide `sup` explicitly. -/ +def conditionallyCompleteLatticeOfLatticeOfsInf (α : Type*) [H1 : Lattice α] [InfSet α] + (isGLB_sInf : ∀ s : Set α, BddBelow s → s.Nonempty → IsGLB s (sInf s)) : + ConditionallyCompleteLattice α := + { H1, + conditionallyCompleteLatticeOfsInf α + (fun a b => ⟨a ⊔ b, forall_insert_of_forall (forall_eq.mpr le_sup_right) le_sup_left⟩) + (fun a b => ⟨a ⊓ b, forall_insert_of_forall (forall_eq.mpr inf_le_right) inf_le_left⟩) + isGLB_sInf with } diff --git a/scripts/noshake.json b/scripts/noshake.json index 86cfab061a740..4a06c0c7680d4 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -385,6 +385,8 @@ ["Mathlib.Control.Basic", "Mathlib.Data.Option.Basic"], "Mathlib.Data.LazyList.Basic": ["Mathlib.Lean.Thunk"], "Mathlib.Data.Int.Defs": ["Batteries.Data.Int.Order"], + "Mathlib.Data.Int.ConditionallyCompleteOrder": + ["Mathlib.Order.ConditionallyCompleteLattice.Basic"], "Mathlib.Data.FunLike.Basic": ["Mathlib.Logic.Function.Basic"], "Mathlib.Data.Finset.Insert": ["Mathlib.Data.Finset.Attr"], "Mathlib.Data.ENat.Lattice": ["Mathlib.Algebra.Group.Action.Defs"], From c4f307606793b52847c977499c58d745bc2186d3 Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 25 Nov 2024 11:04:41 +0000 Subject: [PATCH 279/829] chore: remove unused variables (#19413) #17715 --- Mathlib/Algebra/Group/Pointwise/Set/Card.lean | 2 +- Mathlib/Algebra/GroupWithZero/Pointwise/Set/Card.lean | 2 +- Mathlib/Algebra/Homology/Additive.lean | 4 ++-- Mathlib/Algebra/Quaternion.lean | 4 ++-- Mathlib/Algebra/TrivSqZeroExt.lean | 5 ++--- Mathlib/AlgebraicGeometry/Morphisms/Basic.lean | 5 ++--- Mathlib/AlgebraicGeometry/Morphisms/Etale.lean | 2 +- Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean | 2 +- Mathlib/AlgebraicGeometry/SpreadingOut.lean | 3 +-- Mathlib/Analysis/Convex/Gauge.lean | 2 +- Mathlib/Analysis/Convex/Normed.lean | 4 ++-- Mathlib/Analysis/InnerProductSpace/Basic.lean | 10 +++------- Mathlib/Analysis/Normed/Field/Basic.lean | 4 ++-- Mathlib/Analysis/Normed/Group/Basic.lean | 8 ++++---- Mathlib/Analysis/Normed/Group/Bounded.lean | 3 +-- Mathlib/Analysis/Normed/Group/Uniform.lean | 8 ++++---- Mathlib/Analysis/Normed/Module/FiniteDimension.lean | 7 ++----- Mathlib/CategoryTheory/Adjunction/Evaluation.lean | 1 - Mathlib/CategoryTheory/GradedObject.lean | 2 +- Mathlib/CategoryTheory/GuitartExact/Basic.lean | 1 - .../CategoryTheory/Sites/NonabelianCohomology/H1.lean | 2 +- Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean | 3 +-- .../Combinatorics/SimpleGraph/Triangle/Counting.lean | 4 ++-- .../Combinatorics/SimpleGraph/Triangle/Tripartite.lean | 2 +- Mathlib/Data/Real/Sqrt.lean | 2 -- Mathlib/MeasureTheory/Constructions/Pi.lean | 2 +- Mathlib/MeasureTheory/Group/Action.lean | 4 ++-- Mathlib/MeasureTheory/Group/LIntegral.lean | 2 +- Mathlib/MeasureTheory/Integral/Marginal.lean | 2 +- Mathlib/MeasureTheory/MeasurableSpace/Basic.lean | 4 ++-- Mathlib/MeasureTheory/MeasurableSpace/Prod.lean | 2 +- Mathlib/MeasureTheory/Measure/Regular.lean | 4 ++-- .../GradedAlgebra/HomogeneousLocalization.lean | 3 +-- Mathlib/RingTheory/Localization/AtPrime.lean | 2 +- Mathlib/RingTheory/Localization/Integral.lean | 2 +- .../Localization/LocalizationLocalization.lean | 5 ++--- Mathlib/RingTheory/TensorProduct/Free.lean | 2 +- Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean | 2 -- Mathlib/Testing/Plausible/Functions.lean | 4 ++-- Mathlib/Topology/Algebra/InfiniteSum/Basic.lean | 6 +++--- .../Topology/Algebra/InfiniteSum/Constructions.lean | 2 +- Mathlib/Topology/Algebra/InfiniteSum/Defs.lean | 2 +- Mathlib/Topology/Algebra/InfiniteSum/Group.lean | 4 ++-- Mathlib/Topology/Algebra/InfiniteSum/Ring.lean | 9 ++++----- Mathlib/Topology/ContinuousMap/Ordered.lean | 3 +-- Mathlib/Topology/Instances/ENNReal.lean | 2 +- Mathlib/Topology/MetricSpace/Dilation.lean | 6 +++--- Mathlib/Topology/MetricSpace/Isometry.lean | 2 +- Mathlib/Topology/MetricSpace/Pseudo/Real.lean | 4 ++-- Mathlib/Topology/Order/ExtrClosure.lean | 2 +- Mathlib/Topology/Order/MonotoneConvergence.lean | 8 ++++---- Mathlib/Topology/Order/T5.lean | 3 +-- Mathlib/Topology/ShrinkingLemma.lean | 2 +- Mathlib/Topology/UniformSpace/OfCompactT2.lean | 2 +- 54 files changed, 83 insertions(+), 106 deletions(-) diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean index 9472500e7fcc9..404fdd46e25ad 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean @@ -38,7 +38,7 @@ lemma natCard_mul_le : Nat.card (s * t) ≤ Nat.card s * Nat.card t := by end Mul section InvolutiveInv -variable [InvolutiveInv G] {s t : Set G} +variable [InvolutiveInv G] @[to_additive (attr := simp)] lemma _root_.Cardinal.mk_inv (s : Set G) : #↥(s⁻¹) = #s := by diff --git a/Mathlib/Algebra/GroupWithZero/Pointwise/Set/Card.lean b/Mathlib/Algebra/GroupWithZero/Pointwise/Set/Card.lean index 8466c8876f1c1..4e144b113a816 100644 --- a/Mathlib/Algebra/GroupWithZero/Pointwise/Set/Card.lean +++ b/Mathlib/Algebra/GroupWithZero/Pointwise/Set/Card.lean @@ -13,7 +13,7 @@ import Mathlib.SetTheory.Cardinal.Finite open scoped Cardinal Pointwise -variable {G G₀ M M₀ : Type*} +variable {G₀ M₀ : Type*} namespace Set variable [GroupWithZero G₀] [Zero M₀] [MulActionWithZero G₀ M₀] {a : G₀} diff --git a/Mathlib/Algebra/Homology/Additive.lean b/Mathlib/Algebra/Homology/Additive.lean index 5fbebb2c8beb9..1b14c45bf5b00 100644 --- a/Mathlib/Algebra/Homology/Additive.lean +++ b/Mathlib/Algebra/Homology/Additive.lean @@ -23,8 +23,8 @@ variable {ι : Type*} variable {V : Type u} [Category.{v} V] [Preadditive V] variable {W : Type*} [Category W] [Preadditive W] variable {W₁ W₂ : Type*} [Category W₁] [Category W₂] [HasZeroMorphisms W₁] [HasZeroMorphisms W₂] -variable {c : ComplexShape ι} {C D E : HomologicalComplex V c} -variable (f g : C ⟶ D) (h k : D ⟶ E) (i : ι) +variable {c : ComplexShape ι} {C D : HomologicalComplex V c} +variable (f : C ⟶ D) (i : ι) namespace HomologicalComplex diff --git a/Mathlib/Algebra/Quaternion.lean b/Mathlib/Algebra/Quaternion.lean index 69ce497aff36a..1ee91c3592cc2 100644 --- a/Mathlib/Algebra/Quaternion.lean +++ b/Mathlib/Algebra/Quaternion.lean @@ -93,7 +93,7 @@ theorem equivTuple_apply {R : Type*} (c₁ c₂ : R) (x : ℍ[R,c₁,c₂]) : @[simp] theorem mk.eta {R : Type*} {c₁ c₂} (a : ℍ[R,c₁,c₂]) : mk a.1 a.2 a.3 a.4 = a := rfl -variable {S T R : Type*} {c₁ c₂ : R} (r x y z : R) (a b c : ℍ[R,c₁,c₂]) +variable {S T R : Type*} {c₁ c₂ : R} (r x y : R) (a b : ℍ[R,c₁,c₂]) instance [Subsingleton R] : Subsingleton ℍ[R, c₁, c₂] := (equivTuple c₁ c₂).subsingleton instance [Nontrivial R] : Nontrivial ℍ[R, c₁, c₂] := (equivTuple c₁ c₂).surjective.nontrivial @@ -742,7 +742,7 @@ instance {R : Type*} [One R] [Neg R] [Nontrivial R] : Nontrivial ℍ[R] := namespace Quaternion -variable {S T R : Type*} [CommRing R] (r x y z : R) (a b c : ℍ[R]) +variable {S T R : Type*} [CommRing R] (r x y : R) (a b : ℍ[R]) export QuaternionAlgebra (re imI imJ imK) diff --git a/Mathlib/Algebra/TrivSqZeroExt.lean b/Mathlib/Algebra/TrivSqZeroExt.lean index c40ab975c18d7..cc4239ab85e25 100644 --- a/Mathlib/Algebra/TrivSqZeroExt.lean +++ b/Mathlib/Algebra/TrivSqZeroExt.lean @@ -870,10 +870,9 @@ section Algebra variable (S : Type*) (R R' : Type u) (M : Type v) variable [CommSemiring S] [Semiring R] [CommSemiring R'] [AddCommMonoid M] -variable [Algebra S R] [Algebra S R'] [Module S M] -variable [Module R M] [Module Rᵐᵒᵖ M] [SMulCommClass R Rᵐᵒᵖ M] +variable [Algebra S R] [Module S M] [Module R M] [Module Rᵐᵒᵖ M] [SMulCommClass R Rᵐᵒᵖ M] variable [IsScalarTower S R M] [IsScalarTower S Rᵐᵒᵖ M] -variable [Module R' M] [Module R'ᵐᵒᵖ M] [IsCentralScalar R' M] [IsScalarTower S R' M] +variable [Module R' M] [Module R'ᵐᵒᵖ M] [IsCentralScalar R' M] instance algebra' : Algebra S (tsze R M) := { (TrivSqZeroExt.inlHom R M).comp (algebraMap S R) with diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean b/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean index 51f9c3ccf41d9..070c60d31cc77 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean @@ -142,8 +142,7 @@ instance inf (P Q : MorphismProperty Scheme) [IsLocalAtTarget P] [IsLocalAtTarge fun h ↦ ⟨(iff_of_openCover' f 𝒰).mpr (fun i ↦ (h i).left), (iff_of_openCover' f 𝒰).mpr (fun i ↦ (h i).right)⟩⟩ -variable {P} [hP : IsLocalAtTarget P] -variable {X Y U V : Scheme.{u}} {f : X ⟶ Y} {g : U ⟶ Y} [IsOpenImmersion g] (𝒰 : Y.OpenCover) +variable {P} [hP : IsLocalAtTarget P] {X Y : Scheme.{u}} {f : X ⟶ Y} (𝒰 : Y.OpenCover) lemma of_isPullback {UX UY : Scheme.{u}} {iY : UY ⟶ Y} [IsOpenImmersion iY] {iX : UX ⟶ X} {f' : UX ⟶ UY} (h : IsPullback iX f' f iY) (H : P f) : P f' := by @@ -228,7 +227,7 @@ instance inf (P Q : MorphismProperty Scheme) [IsLocalAtSource P] [IsLocalAtSourc (iff_of_openCover' f 𝒰).mpr (fun i ↦ (h i).right)⟩⟩ variable {P} [IsLocalAtSource P] -variable {X Y U V : Scheme.{u}} {f : X ⟶ Y} {g : U ⟶ Y} [IsOpenImmersion g] (𝒰 : X.OpenCover) +variable {X Y : Scheme.{u}} {f : X ⟶ Y} (𝒰 : X.OpenCover) lemma comp {UX : Scheme.{u}} (H : P f) (i : UX ⟶ X) [IsOpenImmersion i] : P (i ≫ f) := diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Etale.lean b/Mathlib/AlgebraicGeometry/Morphisms/Etale.lean index 665513b4b09cb..35ee3280f7055 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Etale.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Etale.lean @@ -28,7 +28,7 @@ abbrev IsEtale {X Y : Scheme.{u}} (f : X ⟶ Y) := IsSmoothOfRelativeDimension 0 namespace IsEtale -variable {X Y : Scheme.{u}} (f : X ⟶ Y) +variable {X : Scheme.{u}} instance : IsStableUnderBaseChange @IsEtale := isSmoothOfRelativeDimension_isStableUnderBaseChange 0 diff --git a/Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean index 880a1a62832f1..6cd62e8973f1b 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean @@ -26,7 +26,7 @@ universe u namespace AlgebraicGeometry -variable {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) +variable {X Y : Scheme.{u}} theorem isOpenImmersion_iff_stalk {f : X ⟶ Y} : IsOpenImmersion f ↔ IsOpenEmbedding f.base ∧ ∀ x, IsIso (f.stalkMap x) := by diff --git a/Mathlib/AlgebraicGeometry/SpreadingOut.lean b/Mathlib/AlgebraicGeometry/SpreadingOut.lean index b5d7ef83b97b8..a327e1ae7ee02 100644 --- a/Mathlib/AlgebraicGeometry/SpreadingOut.lean +++ b/Mathlib/AlgebraicGeometry/SpreadingOut.lean @@ -49,8 +49,7 @@ open CategoryTheory namespace AlgebraicGeometry -variable {X Y S : Scheme.{u}} (f : X ⟶ Y) (sX : X ⟶ S) (sY : Y ⟶ S) (e : f ≫ sY = sX) -variable {R A : CommRingCat.{u}} +variable {X Y S : Scheme.{u}} (f : X ⟶ Y) (sX : X ⟶ S) (sY : Y ⟶ S) {R A : CommRingCat.{u}} /-- The germ map at `x` is injective if there exists some affine `U ∋ x` such that the map `Γ(X, U) ⟶ X_x` is injective -/ diff --git a/Mathlib/Analysis/Convex/Gauge.lean b/Mathlib/Analysis/Convex/Gauge.lean index 5ecb5e8b2c4cd..f9d037d23877a 100644 --- a/Mathlib/Analysis/Convex/Gauge.lean +++ b/Mathlib/Analysis/Convex/Gauge.lean @@ -41,7 +41,7 @@ open scoped Pointwise Topology NNReal noncomputable section -variable {𝕜 E F : Type*} +variable {𝕜 E : Type*} section AddCommGroup diff --git a/Mathlib/Analysis/Convex/Normed.lean b/Mathlib/Analysis/Convex/Normed.lean index fae3e58b55f18..9a305c1457de0 100644 --- a/Mathlib/Analysis/Convex/Normed.lean +++ b/Mathlib/Analysis/Convex/Normed.lean @@ -25,14 +25,14 @@ We prove the following facts: is bounded. -/ -variable {ι : Type*} {E P : Type*} +variable {E P : Type*} open AffineBasis Module Metric Set open scoped Convex Pointwise Topology section SeminormedAddCommGroup variable [SeminormedAddCommGroup E] [NormedSpace ℝ E] [PseudoMetricSpace P] [NormedAddTorsor E P] -variable {s t : Set E} +variable {s : Set E} /-- The norm on a real normed space is convex on any convex set. See also `Seminorm.convexOn` and `convexOn_univ_norm`. -/ diff --git a/Mathlib/Analysis/InnerProductSpace/Basic.lean b/Mathlib/Analysis/InnerProductSpace/Basic.lean index 2f715eb17ed19..6f06b27cb078c 100644 --- a/Mathlib/Analysis/InnerProductSpace/Basic.lean +++ b/Mathlib/Analysis/InnerProductSpace/Basic.lean @@ -1600,9 +1600,7 @@ end ContinuousLinearMap section -variable {ι : Type*} {ι' : Type*} {ι'' : Type*} -variable {E' : Type*} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] -variable {E'' : Type*} [SeminormedAddCommGroup E''] [InnerProductSpace 𝕜 E''] +variable {ι : Type*} {ι' : Type*} {E' : Type*} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] @[simp] theorem Orthonormal.equiv_refl {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) : @@ -1683,7 +1681,7 @@ open scoped InnerProductSpace variable [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] variable [NormedAddCommGroup F] [InnerProductSpace ℝ F] -variable {ι : Type*} {ι' : Type*} {ι'' : Type*} +variable {ι : Type*} local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y @@ -2188,9 +2186,7 @@ local notation "IK" => @RCLike.I 𝕜 _ local postfix:90 "†" => starRingEnd _ -variable {ι : Type*} -variable {G : ι → Type*} [∀ i, NormedAddCommGroup (G i)] [∀ i, InnerProductSpace 𝕜 (G i)] - {V : ∀ i, G i →ₗᵢ[𝕜] E} (hV : OrthogonalFamily 𝕜 G V) [dec_V : ∀ (i) (x : G i), Decidable (x ≠ 0)] +variable {ι : Type*} {G : ι → Type*} /-- An orthogonal family forms an independent family of subspaces; that is, any collection of elements each from a different subspace in the family is linearly independent. In particular, the diff --git a/Mathlib/Analysis/Normed/Field/Basic.lean b/Mathlib/Analysis/Normed/Field/Basic.lean index eb3c22b61051f..00424d49f8973 100644 --- a/Mathlib/Analysis/Normed/Field/Basic.lean +++ b/Mathlib/Analysis/Normed/Field/Basic.lean @@ -955,7 +955,7 @@ theorem NormedAddCommGroup.tendsto_atTop' [Nonempty α] [Preorder α] [IsDirecte section RingHomIsometric -variable {R₁ : Type*} {R₂ : Type*} {R₃ : Type*} +variable {R₁ R₂ : Type*} /-- This class states that a ring homomorphism is isometric. This is a sufficient assumption for a continuous semilinear map to be bounded and this is the main use for this typeclass. -/ @@ -965,7 +965,7 @@ class RingHomIsometric [Semiring R₁] [Semiring R₂] [Norm R₁] [Norm R₂] ( attribute [simp] RingHomIsometric.is_iso -variable [SeminormedRing R₁] [SeminormedRing R₂] [SeminormedRing R₃] +variable [SeminormedRing R₁] instance RingHomIsometric.ids : RingHomIsometric (RingHom.id R₁) := ⟨rfl⟩ diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index 91cdbe19d79a5..dd3bf2a5b2797 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -44,7 +44,7 @@ normed group -/ -variable {𝓕 𝕜 α ι κ E F G : Type*} +variable {𝓕 α ι κ E F G : Type*} open Filter Function Metric Bornology @@ -337,7 +337,7 @@ abbrev GroupNorm.toNormedCommGroup [CommGroup E] (f : GroupNorm E) : NormedCommG section SeminormedGroup variable [SeminormedGroup E] [SeminormedGroup F] [SeminormedGroup G] {s : Set E} - {a a₁ a₂ b b₁ b₂ c : E} {r r₁ r₂ : ℝ} + {a a₁ a₂ b c : E} {r r₁ r₂ : ℝ} @[to_additive] theorem dist_eq_norm_div (a b : E) : dist a b = ‖a / b‖ := @@ -1024,7 +1024,7 @@ end Induced section SeminormedCommGroup -variable [SeminormedCommGroup E] [SeminormedCommGroup F] {a a₁ a₂ b b₁ b₂ : E} {r r₁ r₂ : ℝ} +variable [SeminormedCommGroup E] [SeminormedCommGroup F] {a b : E} {r : ℝ} @[to_additive] theorem dist_inv (x y : E) : dist x⁻¹ y = dist x y⁻¹ := by @@ -1290,7 +1290,7 @@ end SeminormedCommGroup section NormedGroup -variable [NormedGroup E] [NormedGroup F] {a b : E} +variable [NormedGroup E] {a b : E} @[to_additive (attr := simp) norm_eq_zero] theorem norm_eq_zero'' : ‖a‖ = 0 ↔ a = 1 := diff --git a/Mathlib/Analysis/Normed/Group/Bounded.lean b/Mathlib/Analysis/Normed/Group/Bounded.lean index bd28a0b52f340..dcebdaec8bfb2 100644 --- a/Mathlib/Analysis/Normed/Group/Bounded.lean +++ b/Mathlib/Analysis/Normed/Group/Bounded.lean @@ -21,11 +21,10 @@ normed group open Filter Metric Bornology open scoped Pointwise Topology -variable {α ι E F G : Type*} +variable {α E F G : Type*} section SeminormedGroup variable [SeminormedGroup E] [SeminormedGroup F] [SeminormedGroup G] {s : Set E} - {a a₁ a₂ b b₁ b₂ : E} {r r₁ r₂ : ℝ} @[to_additive (attr := simp) comap_norm_atTop] lemma comap_norm_atTop' : comap norm atTop = cobounded E := by diff --git a/Mathlib/Analysis/Normed/Group/Uniform.lean b/Mathlib/Analysis/Normed/Group/Uniform.lean index 8f4a7f9f06b4e..73b33458e9210 100644 --- a/Mathlib/Analysis/Normed/Group/Uniform.lean +++ b/Mathlib/Analysis/Normed/Group/Uniform.lean @@ -15,13 +15,13 @@ This file proves lipschitzness of normed group operations and shows that normed groups. -/ -variable {𝓕 α E F : Type*} +variable {𝓕 E F : Type*} open Filter Function Metric Bornology open scoped ENNReal NNReal Uniformity Pointwise Topology section SeminormedGroup -variable [SeminormedGroup E] [SeminormedGroup F] {s : Set E} {a a₁ a₂ b b₁ b₂ : E} {r r₁ r₂ : ℝ} +variable [SeminormedGroup E] [SeminormedGroup F] {s : Set E} {a b : E} {r : ℝ} @[to_additive] instance NormedGroup.to_isometricSMul_right : IsometricSMul Eᵐᵒᵖ E := @@ -177,7 +177,7 @@ end SeminormedGroup section SeminormedCommGroup -variable [SeminormedCommGroup E] [SeminormedCommGroup F] {a a₁ a₂ b b₁ b₂ : E} {r r₁ r₂ : ℝ} +variable [SeminormedCommGroup E] [SeminormedCommGroup F] {a₁ a₂ b₁ b₂ : E} {r₁ r₂ : ℝ} @[to_additive] instance NormedGroup.to_isometricSMul_left : IsometricSMul E E := @@ -239,7 +239,7 @@ theorem edist_mul_mul_le (a₁ a₂ b₁ b₂ : E) : section PseudoEMetricSpace variable {α E : Type*} [SeminormedCommGroup E] [PseudoEMetricSpace α] {K Kf Kg : ℝ≥0} - {f g : α → E} {s : Set α} {x : α} + {f g : α → E} {s : Set α} @[to_additive (attr := simp)] lemma lipschitzWith_inv_iff : LipschitzWith K f⁻¹ ↔ LipschitzWith K f := by simp [LipschitzWith] diff --git a/Mathlib/Analysis/Normed/Module/FiniteDimension.lean b/Mathlib/Analysis/Normed/Module/FiniteDimension.lean index 303247548cea2..789c690f2042f 100644 --- a/Mathlib/Analysis/Normed/Module/FiniteDimension.lean +++ b/Mathlib/Analysis/Normed/Module/FiniteDimension.lean @@ -54,8 +54,7 @@ namespace LinearIsometry open LinearMap -variable {R : Type*} [Semiring R] -variable {F E₁ : Type*} [SeminormedAddCommGroup F] [NormedAddCommGroup E₁] [Module R E₁] +variable {F E₁ : Type*} [SeminormedAddCommGroup F] [NormedAddCommGroup E₁] variable {R₁ : Type*} [Field R₁] [Module R₁ E₁] [Module R₁ F] [FiniteDimensional R₁ E₁] [FiniteDimensional R₁ F] @@ -110,9 +109,7 @@ end AffineIsometry section CompleteField variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] {E : Type v} [NormedAddCommGroup E] - [NormedSpace 𝕜 E] {F : Type w} [NormedAddCommGroup F] [NormedSpace 𝕜 F] {F' : Type x} - [AddCommGroup F'] [Module 𝕜 F'] [TopologicalSpace F'] [TopologicalAddGroup F'] - [ContinuousSMul 𝕜 F'] [CompleteSpace 𝕜] + [NormedSpace 𝕜 E] {F : Type w} [NormedAddCommGroup F] [NormedSpace 𝕜 F] [CompleteSpace 𝕜] section Affine diff --git a/Mathlib/CategoryTheory/Adjunction/Evaluation.lean b/Mathlib/CategoryTheory/Adjunction/Evaluation.lean index d11a3d4ac62c1..ffb59cad14a86 100644 --- a/Mathlib/CategoryTheory/Adjunction/Evaluation.lean +++ b/Mathlib/CategoryTheory/Adjunction/Evaluation.lean @@ -22,7 +22,6 @@ open CategoryTheory.Limits universe v₁ v₂ v₃ u₁ u₂ u₃ variable {C : Type u₁} [Category.{v₁} C] (D : Type u₂) [Category.{v₂} D] - (E : Type u₃) [Category.{v₃} E] noncomputable section diff --git a/Mathlib/CategoryTheory/GradedObject.lean b/Mathlib/CategoryTheory/GradedObject.lean index b6e2df434c2b9..2bc160aeba95d 100644 --- a/Mathlib/CategoryTheory/GradedObject.lean +++ b/Mathlib/CategoryTheory/GradedObject.lean @@ -316,7 +316,7 @@ lemma hasMap_of_iso (e : X ≅ Y) (p: I → J) [HasMap X p] : HasMap Y p := fun exact hasColimitOfIso α.symm section -variable [X.HasMap p] [Y.HasMap p] [Z.HasMap p] +variable [X.HasMap p] [Y.HasMap p] /-- Given `X : GradedObject I C` and `p : I → J`, `X.mapObj p` is the graded object by `J` which in degree `j` consists of the coproduct of the `X i` such that `p i = j`. -/ diff --git a/Mathlib/CategoryTheory/GuitartExact/Basic.lean b/Mathlib/CategoryTheory/GuitartExact/Basic.lean index cb64500850d99..b697c1ee0f1a1 100644 --- a/Mathlib/CategoryTheory/GuitartExact/Basic.lean +++ b/Mathlib/CategoryTheory/GuitartExact/Basic.lean @@ -128,7 +128,6 @@ abbrev CostructuredArrowDownwards.mk (comm : R.map a ≫ w.app X₁ ≫ B.map b CostructuredArrow.mk (Y := StructuredArrow.mk a) (StructuredArrow.homMk b (by simpa using comm)) -variable (comm : R.map a ≫ w.app X₁ ≫ B.map b = g) variable {w g} lemma StructuredArrowRightwards.mk_surjective diff --git a/Mathlib/CategoryTheory/Sites/NonabelianCohomology/H1.lean b/Mathlib/CategoryTheory/Sites/NonabelianCohomology/H1.lean index 44a948a9222b1..b15c8b4f9117f 100644 --- a/Mathlib/CategoryTheory/Sites/NonabelianCohomology/H1.lean +++ b/Mathlib/CategoryTheory/Sites/NonabelianCohomology/H1.lean @@ -48,7 +48,7 @@ variable {C : Type u} [Category.{v} C] namespace PresheafOfGroups -variable (G : Cᵒᵖ ⥤ Grp.{w}) {X : C} {I : Type w'} (U : I → C) +variable (G : Cᵒᵖ ⥤ Grp.{w}) {I : Type w'} (U : I → C) /-- A zero cochain consists of a family of sections. -/ def ZeroCochain := ∀ (i : I), G.obj (Opposite.op (U i)) diff --git a/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean b/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean index 6d675a7865092..2924a54b4d530 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean @@ -34,8 +34,7 @@ open Fintype (card) namespace SimpleGraph -variable {α β 𝕜 : Type*} [LinearOrderedField 𝕜] {G H : SimpleGraph α} {ε δ : 𝕜} {n : ℕ} - {s : Finset α} +variable {α β 𝕜 : Type*} [LinearOrderedField 𝕜] {G H : SimpleGraph α} {ε δ : 𝕜} section LocallyLinear diff --git a/Mathlib/Combinatorics/SimpleGraph/Triangle/Counting.lean b/Mathlib/Combinatorics/SimpleGraph/Triangle/Counting.lean index bff996c3fa38c..3de98cd5194b3 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Triangle/Counting.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Triangle/Counting.lean @@ -23,7 +23,7 @@ attribute [-instance] decidableEq_of_subsingleton open Finset Fintype -variable {α : Type*} (G G' : SimpleGraph α) [DecidableRel G.Adj] {ε : ℝ} {s t u : Finset α} +variable {α : Type*} (G : SimpleGraph α) [DecidableRel G.Adj] {ε : ℝ} {s t u : Finset α} namespace SimpleGraph @@ -146,7 +146,7 @@ private lemma triple_eq_triple_of_mem (hst : Disjoint s t) (hsu : Disjoint s u) · rintro rfl solve_by_elim -variable [Fintype α] {P : Finpartition (univ : Finset α)} +variable [Fintype α] /-- The **Triangle Counting Lemma**. If `G` is a graph and `s`, `t`, `u` are disjoint sets of vertices such that each pair is `ε`-uniform and `2 * ε`-dense, then `G` contains at least diff --git a/Mathlib/Combinatorics/SimpleGraph/Triangle/Tripartite.lean b/Mathlib/Combinatorics/SimpleGraph/Triangle/Tripartite.lean index 3135218de5e3d..c911166bc53ac 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Triangle/Tripartite.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Triangle/Tripartite.lean @@ -38,7 +38,7 @@ This construction shows up unrelatedly twice in the theory of Roth numbers: open Finset Function Sum3 variable {α β γ 𝕜 : Type*} [LinearOrderedField 𝕜] {t : Finset (α × β × γ)} {a a' : α} {b b' : β} - {c c' : γ} {x : α × β × γ} {ε : 𝕜} + {c c' : γ} {x : α × β × γ} namespace SimpleGraph namespace TripartiteFromTriangles diff --git a/Mathlib/Data/Real/Sqrt.lean b/Mathlib/Data/Real/Sqrt.lean index b43bd46488f69..5b1193b70b404 100644 --- a/Mathlib/Data/Real/Sqrt.lean +++ b/Mathlib/Data/Real/Sqrt.lean @@ -311,8 +311,6 @@ end Mathlib.Meta.Positivity namespace Real -variable {x y : ℝ} - @[simp] theorem sqrt_mul {x : ℝ} (hx : 0 ≤ x) (y : ℝ) : √(x * y) = √x * √y := by simp_rw [Real.sqrt, ← NNReal.coe_mul, NNReal.coe_inj, Real.toNNReal_mul hx, NNReal.sqrt_mul] diff --git a/Mathlib/MeasureTheory/Constructions/Pi.lean b/Mathlib/MeasureTheory/Constructions/Pi.lean index aafd0f1b3c5ef..670e4e5bf0347 100644 --- a/Mathlib/MeasureTheory/Constructions/Pi.lean +++ b/Mathlib/MeasureTheory/Constructions/Pi.lean @@ -80,7 +80,7 @@ theorem isPiSystem_pi [∀ i, MeasurableSpace (α i)] : section Finite -variable [Finite ι] [Finite ι'] +variable [Finite ι] /-- Boxes of countably spanning sets are countably spanning. -/ theorem IsCountablySpanning.pi {C : ∀ i, Set (Set (α i))} (hC : ∀ i, IsCountablySpanning (C i)) : diff --git a/Mathlib/MeasureTheory/Group/Action.lean b/Mathlib/MeasureTheory/Group/Action.lean index 0cb212b5273b1..cad802e003d1f 100644 --- a/Mathlib/MeasureTheory/Group/Action.lean +++ b/Mathlib/MeasureTheory/Group/Action.lean @@ -28,7 +28,7 @@ namespace MeasureTheory universe u v w -variable {G : Type u} {M : Type v} {α : Type w} {s : Set α} +variable {G : Type u} {M : Type v} {α : Type w} namespace SMulInvariantMeasure @@ -198,7 +198,7 @@ instance smulInvariantMeasure_map_smul [SMul M α] [SMul N α] [SMulCommClass N end SMulHomClass -variable (G) {m : MeasurableSpace α} [Group G] [MulAction G α] (c : G) (μ : Measure α) +variable (G) {m : MeasurableSpace α} [Group G] [MulAction G α] (μ : Measure α) variable [MeasurableSpace G] [MeasurableSMul G α] in /-- Equivalent definitions of a measure invariant under a multiplicative action of a group. diff --git a/Mathlib/MeasureTheory/Group/LIntegral.lean b/Mathlib/MeasureTheory/Group/LIntegral.lean index e0344acdeed93..74c8453b516af 100644 --- a/Mathlib/MeasureTheory/Group/LIntegral.lean +++ b/Mathlib/MeasureTheory/Group/LIntegral.lean @@ -20,7 +20,7 @@ open Measure TopologicalSpace open scoped ENNReal -variable {G : Type*} [MeasurableSpace G] {μ : Measure G} {g : G} +variable {G : Type*} [MeasurableSpace G] {μ : Measure G} section MeasurableMul diff --git a/Mathlib/MeasureTheory/Integral/Marginal.lean b/Mathlib/MeasureTheory/Integral/Marginal.lean index 76fcaa682df92..4e19390aa37fb 100644 --- a/Mathlib/MeasureTheory/Integral/Marginal.lean +++ b/Mathlib/MeasureTheory/Integral/Marginal.lean @@ -66,7 +66,7 @@ section LMarginal variable {δ δ' : Type*} {π : δ → Type*} [∀ x, MeasurableSpace (π x)] variable {μ : ∀ i, Measure (π i)} [DecidableEq δ] -variable {s t : Finset δ} {f g : (∀ i, π i) → ℝ≥0∞} {x y : ∀ i, π i} {i : δ} +variable {s t : Finset δ} {f : (∀ i, π i) → ℝ≥0∞} {x : ∀ i, π i} /-- Integrate `f(x₁,…,xₙ)` over all variables `xᵢ` where `i ∈ s`. Return a function in the remaining variables (it will be constant in the `xᵢ` for `i ∈ s`). diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index 4d2eb7fd31ab1..678a6983571d0 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -55,7 +55,7 @@ open Set Encodable Function Equiv Filter MeasureTheory universe uι -variable {α β γ δ δ' : Type*} {ι : Sort uι} {s t u : Set α} +variable {α β γ δ δ' : Type*} {ι : Sort uι} {s : Set α} namespace MeasurableSpace @@ -225,7 +225,7 @@ variable {f g : α → β} section TypeclassMeasurableSpace -variable [MeasurableSpace α] [MeasurableSpace β] [MeasurableSpace γ] +variable [MeasurableSpace α] [MeasurableSpace β] @[nontriviality, measurability] theorem Subsingleton.measurable [Subsingleton α] : Measurable f := fun _ _ => diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Prod.lean b/Mathlib/MeasureTheory/MeasurableSpace/Prod.lean index e9f2db14bfc25..9dec6c3e39e9b 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Prod.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Prod.lean @@ -18,7 +18,7 @@ noncomputable section open Set MeasurableSpace -variable {α β γ : Type*} [MeasurableSpace α] [MeasurableSpace β] [MeasurableSpace γ] +variable {α β : Type*} [MeasurableSpace α] [MeasurableSpace β] /-- The product of generated σ-algebras is the one generated by rectangles, if both generating sets are countably spanning. -/ diff --git a/Mathlib/MeasureTheory/Measure/Regular.lean b/Mathlib/MeasureTheory/Measure/Regular.lean index ed27aa6936522..91b924e0a4f30 100644 --- a/Mathlib/MeasureTheory/Measure/Regular.lean +++ b/Mathlib/MeasureTheory/Measure/Regular.lean @@ -446,7 +446,7 @@ protected theorem FiniteSpanningSetsIn.outerRegular namespace InnerRegularWRT -variable {p q : Set α → Prop} {U s : Set α} {ε r : ℝ≥0∞} +variable {p : Set α → Prop} /-- If the restrictions of a measure to a monotone sequence of sets covering the space are inner regular for some property `p` and all measurable sets, then the measure itself is @@ -637,7 +637,7 @@ end InnerRegularWRT namespace InnerRegular -variable {U : Set α} {ε : ℝ≥0∞} [TopologicalSpace α] +variable [TopologicalSpace α] /-- The measure of a measurable set is the supremum of the measures of compact sets it contains. -/ theorem _root_.MeasurableSet.measure_eq_iSup_isCompact ⦃U : Set α⦄ (hU : MeasurableSet U) diff --git a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean index af6dc235b4928..82a5c54c0a72c 100644 --- a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean +++ b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean @@ -636,8 +636,7 @@ end section mapAway variable [AddCommMonoid ι] [DecidableEq ι] [GradedAlgebra 𝒜] -variable {d e : ι} {f : A} (hf : f ∈ 𝒜 d) {g : A} (hg : g ∈ 𝒜 e) -variable {x : A} (hx : x = f * g) +variable {e : ι} {f : A} {g : A} (hg : g ∈ 𝒜 e) {x : A} (hx : x = f * g) variable (𝒜) diff --git a/Mathlib/RingTheory/Localization/AtPrime.lean b/Mathlib/RingTheory/Localization/AtPrime.lean index 65cccdc214ca9..76d9a9d2cb2e8 100644 --- a/Mathlib/RingTheory/Localization/AtPrime.lean +++ b/Mathlib/RingTheory/Localization/AtPrime.lean @@ -31,7 +31,7 @@ commutative ring, field of fractions -/ -variable {R : Type*} [CommSemiring R] (M : Submonoid R) (S : Type*) [CommSemiring S] +variable {R : Type*} [CommSemiring R] (S : Type*) [CommSemiring S] variable [Algebra R S] {P : Type*} [CommSemiring P] section AtPrime diff --git a/Mathlib/RingTheory/Localization/Integral.lean b/Mathlib/RingTheory/Localization/Integral.lean index 07c55c2708e4b..ad327225f2383 100644 --- a/Mathlib/RingTheory/Localization/Integral.lean +++ b/Mathlib/RingTheory/Localization/Integral.lean @@ -25,7 +25,7 @@ commutative ring, field of fractions variable {R : Type*} [CommRing R] (M : Submonoid R) {S : Type*} [CommRing S] -variable [Algebra R S] {P : Type*} [CommRing P] +variable [Algebra R S] open Polynomial diff --git a/Mathlib/RingTheory/Localization/LocalizationLocalization.lean b/Mathlib/RingTheory/Localization/LocalizationLocalization.lean index 62db6f2a12761..0637b7cd2c470 100644 --- a/Mathlib/RingTheory/Localization/LocalizationLocalization.lean +++ b/Mathlib/RingTheory/Localization/LocalizationLocalization.lean @@ -27,8 +27,7 @@ namespace IsLocalization section LocalizationLocalization -variable {R : Type*} [CommSemiring R] (M : Submonoid R) {S : Type*} [CommSemiring S] -variable [Algebra R S] {P : Type*} [CommSemiring P] +variable {R : Type*} [CommSemiring R] (M : Submonoid R) {S : Type*} [CommSemiring S] [Algebra R S] variable (N : Submonoid S) (T : Type*) [CommSemiring T] [Algebra R T] @@ -251,7 +250,7 @@ end IsLocalization namespace IsFractionRing -variable {R : Type*} [CommRing R] (M : Submonoid R) {S : Type*} [CommRing S] +variable {R : Type*} [CommRing R] (M : Submonoid R) open IsLocalization diff --git a/Mathlib/RingTheory/TensorProduct/Free.lean b/Mathlib/RingTheory/TensorProduct/Free.lean index 1f644e47cf568..a857cfdf874c6 100644 --- a/Mathlib/RingTheory/TensorProduct/Free.lean +++ b/Mathlib/RingTheory/TensorProduct/Free.lean @@ -30,7 +30,7 @@ namespace Algebra namespace TensorProduct -variable {R S A : Type*} +variable {R A : Type*} section Basis diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean index 680aa8f4249fe..4c69bfb655b75 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean @@ -15,8 +15,6 @@ import Mathlib.RingTheory.UniqueFactorizationDomain.NormalizedFactors * `Nat.instUniqueFactorizationMonoid`: the natural numbers have unique factorization -/ -variable {α : Type*} - namespace Nat instance instWfDvdMonoid : WfDvdMonoid ℕ where diff --git a/Mathlib/Testing/Plausible/Functions.lean b/Mathlib/Testing/Plausible/Functions.lean index 9a83323198f47..a9bdf24d8ccb4 100644 --- a/Mathlib/Testing/Plausible/Functions.lean +++ b/Mathlib/Testing/Plausible/Functions.lean @@ -40,9 +40,9 @@ their defining property is invariant through shrinking. Injective functions are an example of how complicated it can get. -/ -universe u v w +universe u v -variable {α : Type u} {β : Type v} {γ : Sort w} +variable {α : Type u} {β : Type v} namespace Plausible diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean b/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean index 74529c1335d72..ba49c204bab72 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean @@ -21,12 +21,12 @@ noncomputable section open Filter Finset Function Topology -variable {α β γ δ : Type*} +variable {α β γ : Type*} section HasProd variable [CommMonoid α] [TopologicalSpace α] -variable {f g : β → α} {a b : α} {s : Finset β} +variable {f g : β → α} {a b : α} /-- Constant one function has product `1` -/ @[to_additive "Constant zero function has sum `0`"] @@ -355,7 +355,7 @@ end HasProd section tprod -variable [CommMonoid α] [TopologicalSpace α] {f g : β → α} {a a₁ a₂ : α} +variable [CommMonoid α] [TopologicalSpace α] {f g : β → α} @[to_additive] theorem tprod_congr_set_coe (f : β → α) {s t : Set β} (h : s = t) : diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean b/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean index 8b1f1a092440e..4386fe90f6bd1 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean @@ -20,7 +20,7 @@ open Filter Finset Function open scoped Topology -variable {α β γ δ : Type*} +variable {α β γ : Type*} /-! ## Product, Sigma and Pi types -/ diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean b/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean index 2bf685cd201e1..cc83507487809 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean @@ -111,7 +111,7 @@ notation3 "∏' "(...)", "r:67:(scoped f => tprod f) => r @[inherit_doc tsum] notation3 "∑' "(...)", "r:67:(scoped f => tsum f) => r -variable {f g : β → α} {a b : α} {s : Finset β} +variable {f : β → α} {a : α} {s : Finset β} @[to_additive] theorem HasProd.multipliable (h : HasProd f a) : Multipliable f := diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Group.lean b/Mathlib/Topology/Algebra/InfiniteSum/Group.lean index 4c5353916aae3..937a415544928 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Group.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Group.lean @@ -20,7 +20,7 @@ open Filter Finset Function open scoped Topology -variable {α β γ δ : Type*} +variable {α β γ : Type*} section TopologicalGroup @@ -197,7 +197,7 @@ theorem multipliable_iff_cauchySeq_finset [CompleteSpace α] {f : β → α} : Multipliable f ↔ CauchySeq fun s : Finset β ↦ ∏ b ∈ s, f b := by classical exact cauchy_map_iff_exists_tendsto.symm -variable [UniformGroup α] {f g : β → α} {a a₁ a₂ : α} +variable [UniformGroup α] {f g : β → α} @[to_additive] theorem cauchySeq_finset_iff_prod_vanishing : diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Ring.lean b/Mathlib/Topology/Algebra/InfiniteSum/Ring.lean index ea55550f54743..aa0cdf8f9a57d 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Ring.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Ring.lean @@ -19,12 +19,12 @@ This file provides lemmas about the interaction between infinite sums and multip open Filter Finset Function -variable {ι κ R α : Type*} +variable {ι κ α : Type*} section NonUnitalNonAssocSemiring -variable [NonUnitalNonAssocSemiring α] [TopologicalSpace α] [TopologicalSemiring α] {f g : ι → α} - {a a₁ a₂ : α} +variable [NonUnitalNonAssocSemiring α] [TopologicalSpace α] [TopologicalSemiring α] {f : ι → α} + {a₁ : α} theorem HasSum.mul_left (a₂) (h : HasSum f a₁) : HasSum (fun i ↦ a₂ * f i) (a₂ * a₁) := by simpa only using h.map (AddMonoidHom.mulLeft a₂) (continuous_const.mul continuous_id) @@ -63,8 +63,7 @@ end NonUnitalNonAssocSemiring section DivisionSemiring -variable [DivisionSemiring α] [TopologicalSpace α] [TopologicalSemiring α] {f g : ι → α} - {a a₁ a₂ : α} +variable [DivisionSemiring α] [TopologicalSpace α] [TopologicalSemiring α] {f : ι → α} {a a₁ a₂ : α} theorem HasSum.div_const (h : HasSum f a) (b : α) : HasSum (fun i ↦ f i / b) (a / b) := by simp only [div_eq_mul_inv, h.mul_right b⁻¹] diff --git a/Mathlib/Topology/ContinuousMap/Ordered.lean b/Mathlib/Topology/ContinuousMap/Ordered.lean index 5ef9876bb5f8c..7677b1c20ea56 100644 --- a/Mathlib/Topology/ContinuousMap/Ordered.lean +++ b/Mathlib/Topology/ContinuousMap/Ordered.lean @@ -13,8 +13,7 @@ import Mathlib.Topology.ContinuousMap.Defs -/ -variable {α : Type*} {β : Type*} {γ : Type*} -variable [TopologicalSpace α] [TopologicalSpace β] [TopologicalSpace γ] +variable {α β : Type*} [TopologicalSpace α] [TopologicalSpace β] namespace ContinuousMap diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index 94b73607b7f71..029ada7f6a78a 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -25,7 +25,7 @@ variable {α : Type*} {β : Type*} {γ : Type*} namespace ENNReal -variable {a b c d : ℝ≥0∞} {r p q : ℝ≥0} {x y z : ℝ≥0∞} {ε ε₁ ε₂ : ℝ≥0∞} {s : Set ℝ≥0∞} +variable {a b : ℝ≥0∞} {r : ℝ≥0} {x : ℝ≥0∞} {ε : ℝ≥0∞} section TopologicalSpace diff --git a/Mathlib/Topology/MetricSpace/Dilation.lean b/Mathlib/Topology/MetricSpace/Dilation.lean index 201c0e06799e3..6ab93147d696b 100644 --- a/Mathlib/Topology/MetricSpace/Dilation.lean +++ b/Mathlib/Topology/MetricSpace/Dilation.lean @@ -74,7 +74,7 @@ end Defs namespace Dilation -variable {α : Type*} {β : Type*} {γ : Type*} {F : Type*} {G : Type*} +variable {α : Type*} {β : Type*} {γ : Type*} {F : Type*} section Setup @@ -229,8 +229,8 @@ end Setup section PseudoEmetricDilation variable [PseudoEMetricSpace α] [PseudoEMetricSpace β] [PseudoEMetricSpace γ] -variable [FunLike F α β] [DilationClass F α β] [FunLike G β γ] [DilationClass G β γ] -variable (f : F) (g : G) {x y z : α} {s : Set α} +variable [FunLike F α β] [DilationClass F α β] +variable (f : F) /-- Every isometry is a dilation of ratio `1`. -/ @[simps] diff --git a/Mathlib/Topology/MetricSpace/Isometry.lean b/Mathlib/Topology/MetricSpace/Isometry.lean index 5e434b37434df..a987f5c09f284 100644 --- a/Mathlib/Topology/MetricSpace/Isometry.lean +++ b/Mathlib/Topology/MetricSpace/Isometry.lean @@ -62,7 +62,7 @@ namespace Isometry section PseudoEmetricIsometry variable [PseudoEMetricSpace α] [PseudoEMetricSpace β] [PseudoEMetricSpace γ] -variable {f : α → β} {x y z : α} {s : Set α} +variable {f : α → β} {x : α} /-- An isometry preserves edistances. -/ theorem edist_eq (hf : Isometry f) (x y : α) : edist (f x) (f y) = edist x y := diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Real.lean b/Mathlib/Topology/MetricSpace/Pseudo/Real.lean index baa4f534f75cc..6183ec868a947 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Real.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Real.lean @@ -15,7 +15,7 @@ open scoped NNReal Topology namespace Real -variable {ι α : Type*} [PseudoMetricSpace α] +variable {ι : Type*} lemma dist_left_le_of_mem_uIcc {x y z : ℝ} (h : y ∈ uIcc x z) : dist x y ≤ dist x z := by simpa only [dist_comm x] using abs_sub_left_of_mem_uIcc h @@ -35,7 +35,7 @@ lemma dist_le_of_mem_Icc {x y x' y' : ℝ} (hx : x ∈ Icc x' y') (hy : y ∈ Ic lemma dist_le_of_mem_Icc_01 {x y : ℝ} (hx : x ∈ Icc (0 : ℝ) 1) (hy : y ∈ Icc (0 : ℝ) 1) : dist x y ≤ 1 := by simpa only [sub_zero] using Real.dist_le_of_mem_Icc hx hy -variable {π : ι → Type*} [Fintype ι] [∀ i, PseudoMetricSpace (π i)] {x y x' y' : ι → ℝ} +variable [Fintype ι] {x y x' y' : ι → ℝ} lemma dist_le_of_mem_pi_Icc (hx : x ∈ Icc x' y') (hy : y ∈ Icc x' y') : dist x y ≤ dist x' y' := by refine (dist_pi_le_iff dist_nonneg).2 fun b => diff --git a/Mathlib/Topology/Order/ExtrClosure.lean b/Mathlib/Topology/Order/ExtrClosure.lean index d1b2049789057..31b06abf781bd 100644 --- a/Mathlib/Topology/Order/ExtrClosure.lean +++ b/Mathlib/Topology/Order/ExtrClosure.lean @@ -20,7 +20,7 @@ open Filter Set open Topology variable {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] [Preorder Y] - [OrderClosedTopology Y] {f g : X → Y} {s : Set X} {a : X} + [OrderClosedTopology Y] {f : X → Y} {s : Set X} {a : X} protected theorem IsMaxOn.closure (h : IsMaxOn f s a) (hc : ContinuousOn f (closure s)) : IsMaxOn f (closure s) a := fun x hx => diff --git a/Mathlib/Topology/Order/MonotoneConvergence.lean b/Mathlib/Topology/Order/MonotoneConvergence.lean index bad7ee58b291e..7936b81bae9f3 100644 --- a/Mathlib/Topology/Order/MonotoneConvergence.lean +++ b/Mathlib/Topology/Order/MonotoneConvergence.lean @@ -107,7 +107,7 @@ end IsGLB section CiSup -variable [ConditionallyCompleteLattice α] [SupConvergenceClass α] {f : ι → α} {a : α} +variable [ConditionallyCompleteLattice α] [SupConvergenceClass α] {f : ι → α} theorem tendsto_atTop_ciSup (h_mono : Monotone f) (hbdd : BddAbove <| range f) : Tendsto f atTop (𝓝 (⨆ i, f i)) := by @@ -121,7 +121,7 @@ end CiSup section CiInf -variable [ConditionallyCompleteLattice α] [InfConvergenceClass α] {f : ι → α} {a : α} +variable [ConditionallyCompleteLattice α] [InfConvergenceClass α] {f : ι → α} theorem tendsto_atBot_ciInf (h_mono : Monotone f) (hbdd : BddBelow <| range f) : Tendsto f atBot (𝓝 (⨅ i, f i)) := by convert tendsto_atTop_ciSup h_mono.dual hbdd.dual using 1 @@ -133,7 +133,7 @@ end CiInf section iSup -variable [CompleteLattice α] [SupConvergenceClass α] {f : ι → α} {a : α} +variable [CompleteLattice α] [SupConvergenceClass α] {f : ι → α} theorem tendsto_atTop_iSup (h_mono : Monotone f) : Tendsto f atTop (𝓝 (⨆ i, f i)) := tendsto_atTop_ciSup h_mono (OrderTop.bddAbove _) @@ -145,7 +145,7 @@ end iSup section iInf -variable [CompleteLattice α] [InfConvergenceClass α] {f : ι → α} {a : α} +variable [CompleteLattice α] [InfConvergenceClass α] {f : ι → α} theorem tendsto_atBot_iInf (h_mono : Monotone f) : Tendsto f atBot (𝓝 (⨅ i, f i)) := tendsto_atBot_ciInf h_mono (OrderBot.bddBelow _) diff --git a/Mathlib/Topology/Order/T5.lean b/Mathlib/Topology/Order/T5.lean index 39e0d08dc9b0a..597085db3f109 100644 --- a/Mathlib/Topology/Order/T5.lean +++ b/Mathlib/Topology/Order/T5.lean @@ -16,8 +16,7 @@ topological space. open Filter Set Function OrderDual Topology Interval -variable {X : Type*} [LinearOrder X] [TopologicalSpace X] [OrderTopology X] {a b c : X} - {s t : Set X} +variable {X : Type*} [LinearOrder X] [TopologicalSpace X] [OrderTopology X] {a : X} {s t : Set X} namespace Set diff --git a/Mathlib/Topology/ShrinkingLemma.lean b/Mathlib/Topology/ShrinkingLemma.lean index 8f03eeab6a068..65f6b72889730 100644 --- a/Mathlib/Topology/ShrinkingLemma.lean +++ b/Mathlib/Topology/ShrinkingLemma.lean @@ -206,7 +206,7 @@ section NormalSpace open ShrinkingLemma -variable {u : ι → Set X} {s : Set X} {p : Set X → Prop} [NormalSpace X] +variable {u : ι → Set X} {s : Set X} [NormalSpace X] /-- **Shrinking lemma**. A point-finite open cover of a closed subset of a normal space can be "shrunk" to a new open cover so that the closure of each new open set is contained in the diff --git a/Mathlib/Topology/UniformSpace/OfCompactT2.lean b/Mathlib/Topology/UniformSpace/OfCompactT2.lean index 5eaf71a852f4f..c9d49e9474570 100644 --- a/Mathlib/Topology/UniformSpace/OfCompactT2.lean +++ b/Mathlib/Topology/UniformSpace/OfCompactT2.lean @@ -26,7 +26,7 @@ uniform space, uniform continuity, compact space open Uniformity Topology Filter UniformSpace Set -variable {α β γ : Type*} [UniformSpace α] [UniformSpace β] +variable {γ : Type*} /-! ### Uniformity on compact spaces From 2fa86db9f80c9a25fe911ff8ba831094a2659078 Mon Sep 17 00:00:00 2001 From: Nick Ward <102917377+gio256@users.noreply.github.com> Date: Mon, 25 Nov 2024 11:42:16 +0000 Subject: [PATCH 280/829] feat(AlgebraicTopology): StrictSegal simplicial sets are quasicategories (#19270) We prove that any simplicial set satisfying the `StrictSegal` condition introduced in #18499 is a quasicategory. This implies that the nerve of a category is a quasicategory. We construct a solution to the lifting problem against `hornInclusion n i` as a `spineToSimplex` of the path in the image of the spine of the standard simplex. To prove that this extension restricts to the appropriate map on the horn, we apply `spineInjective`, reducing the question of equality between two `n`-simplices in `X` to that of equality between the paths along their spines. The remainder of the proof is a straightforward application of basic properties relating spines of simplices to their faces. Co-Authored-By: [Johan Commelin](https://github.com/jcommelin) and [Emily Riehl](https://github.com/emilyriehl) --- .../AlgebraicTopology/SimplexCategory.lean | 44 ++++++ .../SimplicialSet/Basic.lean | 7 + .../AlgebraicTopology/SimplicialSet/Path.lean | 50 ++++++ .../SimplicialSet/StrictSegal.lean | 148 +++++++++++++++++- 4 files changed, 248 insertions(+), 1 deletion(-) diff --git a/Mathlib/AlgebraicTopology/SimplexCategory.lean b/Mathlib/AlgebraicTopology/SimplexCategory.lean index a2e26b4118880..7c9ce0278a0c3 100644 --- a/Mathlib/AlgebraicTopology/SimplexCategory.lean +++ b/Mathlib/AlgebraicTopology/SimplexCategory.lean @@ -564,6 +564,50 @@ lemma δ_one_mkOfSucc {n : ℕ} (i : Fin n) : fin_cases x aesop +/-- If `i + 1 < j`, `mkOfSucc i ≫ δ j` is the morphism `[1] ⟶ [n]` that +sends `0` and `1` to `i` and `i + 1`, respectively. -/ +lemma mkOfSucc_δ_lt {n : ℕ} {i : Fin n} {j : Fin (n + 2)} + (h : i.succ.castSucc < j) : + mkOfSucc i ≫ δ j = mkOfSucc i.castSucc := by + ext x + fin_cases x + · simp [δ, Fin.succAbove_of_castSucc_lt _ _ (Nat.lt_trans _ h)] + · simp [δ, Fin.succAbove_of_castSucc_lt _ _ h] + +/-- If `i + 1 > j`, `mkOfSucc i ≫ δ j` is the morphism `[1] ⟶ [n]` that +sends `0` and `1` to `i + 1` and `i + 2`, respectively. -/ +lemma mkOfSucc_δ_gt {n : ℕ} {i : Fin n} {j : Fin (n + 2)} + (h : j < i.succ.castSucc) : + mkOfSucc i ≫ δ j = mkOfSucc i.succ := by + ext x + simp only [δ, len_mk, mkHom, comp_toOrderHom, Hom.toOrderHom_mk, OrderHom.comp_coe, + OrderEmbedding.toOrderHom_coe, Function.comp_apply, Fin.succAboveOrderEmb_apply] + fin_cases x <;> rw [Fin.succAbove_of_le_castSucc] + · rfl + · exact Nat.le_of_lt_succ h + · rfl + · exact Nat.le_of_lt h + +/-- If `i + 1 = j`, `mkOfSucc i ≫ δ j` is the morphism `[1] ⟶ [n]` that +sends `0` and `1` to `i` and `i + 2`, respectively. -/ +lemma mkOfSucc_δ_eq {n : ℕ} {i : Fin n} {j : Fin (n + 2)} + (h : j = i.succ.castSucc) : + mkOfSucc i ≫ δ j = intervalEdge i 2 (by omega) := by + ext x + fin_cases x + · subst h + simp only [δ, len_mk, Nat.reduceAdd, mkHom, comp_toOrderHom, Hom.toOrderHom_mk, + Fin.zero_eta, OrderHom.comp_coe, OrderEmbedding.toOrderHom_coe, Function.comp_apply, + mkOfSucc_homToOrderHom_zero, Fin.succAboveOrderEmb_apply, + Fin.castSucc_succAbove_castSucc, Fin.succAbove_succ_self] + rfl + · simp only [δ, len_mk, Nat.reduceAdd, mkHom, comp_toOrderHom, Hom.toOrderHom_mk, Fin.mk_one, + OrderHom.comp_coe, OrderEmbedding.toOrderHom_coe, Function.comp_apply, + mkOfSucc_homToOrderHom_one, Fin.succAboveOrderEmb_apply] + subst h + rw [Fin.succAbove_castSucc_self] + rfl + theorem eq_of_one_to_two (f : ([1] : SimplexCategory) ⟶ [2]) : f = (δ (n := 1) 0) ∨ f = (δ (n := 1) 1) ∨ f = (δ (n := 1) 2) ∨ ∃ a, f = SimplexCategory.const _ _ a := by diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean index 309fc78b1916a..8c329680dec88 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean @@ -106,6 +106,13 @@ def _root_.SSet.yonedaEquiv (X : SSet.{u}) (n : SimplexCategory) : (standardSimplex.obj n ⟶ X) ≃ X.obj (op n) := yonedaCompUliftFunctorEquiv X n +/-- The unique non-degenerate `n`-simplex in `Δ[n]`. -/ +def id (n : ℕ) : Δ[n] _[n] := yonedaEquiv Δ[n] [n] (𝟙 Δ[n]) + +lemma id_eq_objEquiv_symm (n : ℕ) : id n = (objEquiv _ _).symm (𝟙 _) := rfl + +lemma objEquiv_id (n : ℕ) : objEquiv _ _ (id n) = 𝟙 _ := rfl + /-- The (degenerate) `m`-simplex in the standard simplex concentrated in vertex `k`. -/ def const (n : ℕ) (k : Fin (n+1)) (m : SimplexCategoryᵒᵖ) : Δ[n].obj m := objMk (OrderHom.const _ k ) diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean index 047e74c18ff93..33e4d07da1bce 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean @@ -85,4 +85,54 @@ lemma Path.ext' {n : ℕ} {f g : Path X (n + 1)} rw [← f.arrow_tgt (Fin.last n), ← g.arrow_tgt (Fin.last n), h] · exact h j +/-- Maps of simplicial sets induce maps of paths in a simplicial set.-/ +@[simps] +def Path.map {X Y : SSet.{u}} {n : ℕ} (f : X.Path n) (σ : X ⟶ Y) : Y.Path n where + vertex i := σ.app (Opposite.op [0]) (f.vertex i) + arrow i := σ.app (Opposite.op [1]) (f.arrow i) + arrow_src i := by + simp only [← f.arrow_src i] + exact congr (σ.naturality (δ 1).op) rfl |>.symm + arrow_tgt i := by + simp only [← f.arrow_tgt i] + exact congr (σ.naturality (δ 0).op) rfl |>.symm + +/-- `Path.map` respects subintervals of paths.-/ +lemma map_interval {X Y : SSet.{u}} {n : ℕ} (f : X.Path n) (σ : X ⟶ Y) + (j l : ℕ) (hjl : j + l ≤ n) : + (f.map σ).interval j l hjl = (f.interval j l hjl).map σ := rfl + +/-- The spine of the unique non-degenerate `n`-simplex in `Δ[n]`.-/ +def standardSimplex.spineId (n : ℕ) : Path Δ[n] n := + spine Δ[n] n (standardSimplex.id n) + +/-- Any inner horn contains the spine of the unique non-degenerate `n`-simplex +in `Δ[n]`.-/ +@[simps] +def horn.spineId {n : ℕ} (i : Fin (n + 3)) + (h₀ : 0 < i) (hₙ : i < Fin.last (n + 2)) : + Path Λ[n + 2, i] (n + 2) where + vertex j := ⟨standardSimplex.spineId _ |>.vertex j, (horn.const n i j _).property⟩ + arrow j := ⟨standardSimplex.spineId _ |>.arrow j, by + let edge := horn.primitiveEdge h₀ hₙ j + have ha : (standardSimplex.spineId _).arrow j = edge.val := by + dsimp only [edge, standardSimplex.spineId, standardSimplex.id, spine_arrow, + mkOfSucc, horn.primitiveEdge, horn.edge, standardSimplex.edge, + standardSimplex.map_apply] + aesop + rw [ha] + exact edge.property⟩ + arrow_src := by + simp only [horn, SimplicialObject.δ, Subtype.mk.injEq] + exact standardSimplex.spineId _ |>.arrow_src + arrow_tgt := by + simp only [horn, SimplicialObject.δ, Subtype.mk.injEq] + exact standardSimplex.spineId _ |>.arrow_tgt + +@[simp] +lemma horn.spineId_map_hornInclusion {n : ℕ} (i : Fin (n + 3)) + (h₀ : 0 < i) (hₙ : i < Fin.last (n + 2)) : + Path.map (horn.spineId i h₀ hₙ) (hornInclusion (n + 2) i) = + standardSimplex.spineId (n + 2) := rfl + end SSet diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean b/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean index 13a41139b88ed..ca5432b7bbe6b 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean @@ -1,10 +1,11 @@ /- Copyright (c) 2024 Emily Riehl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mario Carneiro, Emily Riehl, Joël Riou +Authors: Mario Carneiro, Emily Riehl, Joël Riou, Johan Commelin, Nick Ward -/ import Mathlib.AlgebraicTopology.SimplicialSet.Nerve import Mathlib.AlgebraicTopology.SimplicialSet.Path +import Mathlib.AlgebraicTopology.SimplicialSet.Quasicategory import Mathlib.CategoryTheory.Functor.KanExtension.Adjunction import Mathlib.CategoryTheory.Functor.KanExtension.Basic @@ -86,6 +87,147 @@ theorem spineToSimplex_edge (f : Path X n) (j l : ℕ) (hjl : j + l ≤ n) : unfold diagonal simp only [← FunctorToTypes.map_comp_apply, ← op_comp, diag_subinterval_eq] +/-- For any `σ : X ⟶ Y` between `StrictSegal` simplicial sets, `spineToSimplex` +commutes with `Path.map`. -/ +lemma spineToSimplex_map {X Y : SSet.{u}} [StrictSegal X] [StrictSegal Y] + {n : ℕ} (f : Path X (n + 1)) (σ : X ⟶ Y) : + spineToSimplex (f.map σ) = σ.app _ (spineToSimplex f) := by + apply spineInjective + ext k + dsimp only [spineEquiv, Equiv.coe_fn_mk, Path.map, spine_arrow] + rw [← types_comp_apply (σ.app _) (Y.map _), ← σ.naturality] + simp only [types_comp_apply, spineToSimplex_arrow] + +/-- If we take the path along the spine of the `j`th face of a `spineToSimplex`, +the common vertices will agree with those of the original path `f`. In particular, +a vertex `i` with `i < j` can be identified with the same vertex in `f`. -/ +lemma spine_δ_vertex_lt (f : Path X (n + 1)) {i : Fin (n + 1)} {j : Fin (n + 2)} + (h : i.castSucc < j) : + (X.spine n (X.δ j (spineToSimplex f))).vertex i = f.vertex i.castSucc := by + simp only [SimplicialObject.δ, spine_vertex] + rw [← FunctorToTypes.map_comp_apply, ← op_comp, const_comp, spineToSimplex_vertex] + simp only [SimplexCategory.δ, Hom.toOrderHom, len_mk, mkHom, Hom.mk, + OrderEmbedding.toOrderHom_coe, Fin.succAboveOrderEmb_apply] + rw [Fin.succAbove_of_castSucc_lt j i h] + +/-- If we take the path along the spine of the `j`th face of a `spineToSimplex`, +a vertex `i` with `i ≥ j` can be identified with vertex `i + 1` in the original +path. -/ +lemma spine_δ_vertex_ge (f : Path X (n + 1)) {i : Fin (n + 1)} {j : Fin (n + 2)} + (h : j ≤ i.castSucc) : + (X.spine n (X.δ j (spineToSimplex f))).vertex i = f.vertex i.succ := by + simp only [SimplicialObject.δ, spine_vertex] + rw [← FunctorToTypes.map_comp_apply, ← op_comp, const_comp, spineToSimplex_vertex] + simp only [SimplexCategory.δ, Hom.toOrderHom, len_mk, mkHom, Hom.mk, + OrderEmbedding.toOrderHom_coe, Fin.succAboveOrderEmb_apply] + rw [Fin.succAbove_of_le_castSucc j i h] + +/-- If we take the path along the spine of the `j`th face of a `spineToSimplex`, +the common arrows will agree with those of the original path `f`. In particular, +an arrow `i` with `i + 1 < j` can be identified with the same arrow in `f`. -/ +lemma spine_δ_arrow_lt (f : Path X (n + 1)) {i : Fin n} {j : Fin (n + 2)} + (h : i.succ.castSucc < j) : + (X.spine n (X.δ j (spineToSimplex f))).arrow i = f.arrow i.castSucc := by + simp only [SimplicialObject.δ, spine_arrow] + rw [← FunctorToTypes.map_comp_apply, ← op_comp] + rw [mkOfSucc_δ_lt h, spineToSimplex_arrow] + +/-- If we take the path along the spine of the `j`th face of a `spineToSimplex`, +an arrow `i` with `i + 1 > j` can be identified with arrow `i + 1` in the +original path. -/ +lemma spine_δ_arrow_gt (f : Path X (n + 1)) {i : Fin n} {j : Fin (n + 2)} + (h : j < i.succ.castSucc) : + (X.spine n (X.δ j (spineToSimplex f))).arrow i = f.arrow i.succ := by + simp only [SimplicialObject.δ, spine_arrow] + rw [← FunctorToTypes.map_comp_apply, ← op_comp] + rw [mkOfSucc_δ_gt h, spineToSimplex_arrow] + +/-- If we take the path along the spine of a face of a `spineToSimplex`, the +arrows not contained in the original path can be recovered as the diagonal edge +of the `spineToSimplex` that "composes" arrows `i` and `i + 1`. -/ +lemma spine_δ_arrow_eq (f : Path X (n + 1)) {i : Fin n} {j : Fin (n + 2)} + (h : j = i.succ.castSucc) : + (X.spine n (X.δ j (spineToSimplex f))).arrow i = + spineToDiagonal (Path.interval f i 2 (by omega)) := by + simp only [SimplicialObject.δ, spine_arrow] + rw [← FunctorToTypes.map_comp_apply, ← op_comp] + rw [mkOfSucc_δ_eq h, spineToSimplex_edge] + +/-- Any `StrictSegal` simplicial set is a `Quasicategory`. -/ +instance : Quasicategory X := by + apply quasicategory_of_filler X + intro n i σ₀ h₀ hₙ + use spineToSimplex <| Path.map (horn.spineId i h₀ hₙ) σ₀ + intro j hj + apply spineInjective + ext k + · dsimp only [spineEquiv, spine_arrow, Function.comp_apply, Equiv.coe_fn_mk] + rw [← types_comp_apply (σ₀.app _) (X.map _), ← σ₀.naturality] + let ksucc := k.succ.castSucc + obtain hlt | hgt | heq : ksucc < j ∨ j < ksucc ∨ j = ksucc := by omega + · rw [← spine_arrow, spine_δ_arrow_lt _ hlt] + dsimp only [Path.map, spine_arrow, Fin.coe_eq_castSucc] + apply congr_arg + simp only [horn, horn.spineId, standardSimplex, uliftFunctor, Functor.comp_obj, + yoneda_obj_obj, whiskering_obj_obj_map, uliftFunctor_map, yoneda_obj_map, + standardSimplex.objEquiv, Equiv.ulift, Equiv.coe_fn_symm_mk, + Quiver.Hom.unop_op, horn.face_coe, Subtype.mk.injEq] + rw [mkOfSucc_δ_lt hlt] + rfl + · rw [← spine_arrow, spine_δ_arrow_gt _ hgt] + dsimp only [Path.map, spine_arrow, Fin.coe_eq_castSucc] + apply congr_arg + simp only [horn, horn.spineId, standardSimplex, uliftFunctor, Functor.comp_obj, + yoneda_obj_obj, whiskering_obj_obj_map, uliftFunctor_map, yoneda_obj_map, + standardSimplex.objEquiv, Equiv.ulift, Equiv.coe_fn_symm_mk, + Quiver.Hom.unop_op, horn.face_coe, Subtype.mk.injEq] + rw [mkOfSucc_δ_gt hgt] + rfl + · /- The only inner horn of `Δ[2]` does not contain the diagonal edge. -/ + have hn0 : n ≠ 0 := by + rintro rfl + obtain rfl : k = 0 := by omega + fin_cases i <;> contradiction + /- We construct the triangle in the standard simplex as a 2-simplex in + the horn. While the triangle is not contained in the inner horn `Λ[2, 1]`, + we can inhabit `Λ[n + 2, i] _[2]` by induction on `n`. -/ + let triangle : Λ[n + 2, i] _[2] := by + cases n with + | zero => contradiction + | succ _ => exact horn.primitiveTriangle i h₀ hₙ k (by omega) + /- The interval spanning from `k` to `k + 2` is equivalently the spine + of the triangle with vertices `k`, `k + 1`, and `k + 2`. -/ + have hi : ((horn.spineId i h₀ hₙ).map σ₀).interval k 2 (by omega) = + X.spine 2 (σ₀.app _ triangle) := by + ext m + dsimp [spine_arrow, Path.interval, Path.map] + rw [← types_comp_apply (σ₀.app _) (X.map _), ← σ₀.naturality] + apply congr_arg + simp only [horn, standardSimplex, uliftFunctor, Functor.comp_obj, + whiskering_obj_obj_obj, yoneda_obj_obj, uliftFunctor_obj, ne_eq, + whiskering_obj_obj_map, uliftFunctor_map, yoneda_obj_map, len_mk, + Nat.reduceAdd, Quiver.Hom.unop_op] + cases n with + | zero => contradiction + | succ _ => ext x; fin_cases x <;> fin_cases m <;> rfl + rw [← spine_arrow, spine_δ_arrow_eq _ heq, hi] + simp only [spineToDiagonal, diagonal, spineToSimplex_spine] + rw [← types_comp_apply (σ₀.app _) (X.map _), ← σ₀.naturality, types_comp_apply] + apply congr_arg + simp only [horn, standardSimplex, uliftFunctor, Functor.comp_obj, + whiskering_obj_obj_obj, yoneda_obj_obj, uliftFunctor_obj, + uliftFunctor_map, whiskering_obj_obj_map, yoneda_obj_map, horn.face_coe, + len_mk, Nat.reduceAdd, Quiver.Hom.unop_op, Subtype.mk.injEq, ULift.up_inj] + ext z + cases n with + | zero => contradiction + | succ _ => + fin_cases z <;> + · simp only [standardSimplex.objEquiv, uliftFunctor_map, yoneda_obj_map, + Quiver.Hom.unop_op, Equiv.ulift_symm_down] + rw [mkOfSucc_δ_eq heq] + rfl + end StrictSegal end SSet @@ -118,4 +260,8 @@ noncomputable instance strictSegal (C : Type u) [Category.{v} C] : StrictSegal ( · intro i hi apply ComposableArrows.mkOfObjOfMapSucc_map_succ +/-- By virtue of satisfying the `StrictSegal` condition, the nerve of a +category is a `Quasicategory`. -/ +instance : Quasicategory (nerve C) := inferInstance + end Nerve From 8defd9fe75da7a55a73347ba859b43448550ffb5 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 25 Nov 2024 11:42:18 +0000 Subject: [PATCH 281/829] chore(Filter): move defs of `Filter.IsBounded` etc (#19358) --- Mathlib/Analysis/Normed/Group/Bounded.lean | 1 - Mathlib/Order/Filter/Defs.lean | 35 ++++++++++++++++++++++ Mathlib/Order/LiminfLimsup.lean | 30 ------------------- Mathlib/Topology/Bornology/Basic.lean | 2 +- 4 files changed, 36 insertions(+), 32 deletions(-) diff --git a/Mathlib/Analysis/Normed/Group/Bounded.lean b/Mathlib/Analysis/Normed/Group/Bounded.lean index dcebdaec8bfb2..0795999cb5541 100644 --- a/Mathlib/Analysis/Normed/Group/Bounded.lean +++ b/Mathlib/Analysis/Normed/Group/Bounded.lean @@ -6,7 +6,6 @@ Authors: Patrick Massot, Johannes Hölzl, Yaël Dillies import Mathlib.Analysis.Normed.Group.Basic import Mathlib.Topology.MetricSpace.Bounded import Mathlib.Order.Filter.Pointwise -import Mathlib.Order.LiminfLimsup /-! # Boundedness in normed groups diff --git a/Mathlib/Order/Filter/Defs.lean b/Mathlib/Order/Filter/Defs.lean index f03e4c29d0296..d0a00823afb8a 100644 --- a/Mathlib/Order/Filter/Defs.lean +++ b/Mathlib/Order/Filter/Defs.lean @@ -32,6 +32,11 @@ abstract two related kinds of ideas: a tactic that takes a list of proofs `hᵢ : sᵢ ∈ f`, and replaces a goal `s ∈ f` with `∀ x, x ∈ s₁ → ... → x ∈ sₙ → x ∈ s`; * `Filter.NeBot f` : a utility class stating that `f` is a non-trivial filter. +* `Filter.IsBounded r f`: the filter `f` is eventually bounded w.r.t. the relation `r`, + i.e. eventually, it is bounded by some uniform bound. + `r` will be usually instantiated with `(· ≤ ·)` or `(· ≥ ·)`. +* `Filter.IsCobounded r f` states that the filter `f` does not tend to infinity w.r.t. `r`. + This is also called frequently bounded. Will be usually instantiated with `(· ≤ ·)` or `(· ≥ ·)`. ## Notations @@ -350,6 +355,36 @@ This is essentially a push-forward along a function mapping each set to a set. - protected def lift' (f : Filter α) (h : Set α → Set β) := f.lift (𝓟 ∘ h) +/-- `f.IsBounded r`: the filter `f` is eventually bounded w.r.t. the relation `r`, +i.e. eventually, it is bounded by some uniform bound. +`r` will be usually instantiated with `(· ≤ ·)` or `(· ≥ ·)`. -/ +def IsBounded (r : α → α → Prop) (f : Filter α) := + ∃ b, ∀ᶠ x in f, r x b + +/-- `f.IsBoundedUnder (≺) u`: the image of the filter `f` under `u` is eventually bounded w.r.t. +the relation `≺`, i.e. eventually, it is bounded by some uniform bound. -/ +def IsBoundedUnder (r : α → α → Prop) (f : Filter β) (u : β → α) := + (map u f).IsBounded r + +/-- `IsCobounded (≺) f` states that the filter `f` does not tend to infinity w.r.t. `≺`. This is +also called frequently bounded. Will be usually instantiated with `≤` or `≥`. + +There is a subtlety in this definition: we want `f.IsCobounded` to hold for any `f` in the case of +complete lattices. This will be relevant to deduce theorems on complete lattices from their +versions on conditionally complete lattices with additional assumptions. We have to be careful in +the edge case of the trivial filter containing the empty set: the other natural definition + `¬ ∀ a, ∀ᶠ n in f, a ≤ n` +would not work as well in this case. +-/ +def IsCobounded (r : α → α → Prop) (f : Filter α) := + ∃ b, ∀ a, (∀ᶠ x in f, r x a) → r b a + +/-- `IsCoboundedUnder (≺) f u` states that the image of the filter `f` under the map `u` does not +tend to infinity w.r.t. `≺`. This is also called frequently bounded. Will be usually instantiated +with `≤` or `≥`. -/ +def IsCoboundedUnder (r : α → α → Prop) (f : Filter β) (u : β → α) := + (map u f).IsCobounded r + end Filter namespace Mathlib.Tactic diff --git a/Mathlib/Order/LiminfLimsup.lean b/Mathlib/Order/LiminfLimsup.lean index 2576f34d60217..6cdea2ca4d772 100644 --- a/Mathlib/Order/LiminfLimsup.lean +++ b/Mathlib/Order/LiminfLimsup.lean @@ -46,17 +46,6 @@ namespace Filter section Relation -/-- `f.IsBounded (≺)`: the filter `f` is eventually bounded w.r.t. the relation `≺`, i.e. -eventually, it is bounded by some uniform bound. -`r` will be usually instantiated with `≤` or `≥`. -/ -def IsBounded (r : α → α → Prop) (f : Filter α) := - ∃ b, ∀ᶠ x in f, r x b - -/-- `f.IsBoundedUnder (≺) u`: the image of the filter `f` under `u` is eventually bounded w.r.t. -the relation `≺`, i.e. eventually, it is bounded by some uniform bound. -/ -def IsBoundedUnder (r : α → α → Prop) (f : Filter β) (u : β → α) := - (map u f).IsBounded r - variable {r : α → α → Prop} {f g : Filter α} /-- `f` is eventually bounded if and only if, there exists an admissible set on which it is @@ -218,25 +207,6 @@ theorem IsBoundedUnder.bddBelow_range [Preorder β] [IsDirected β (· ≥ ·)] (hf : IsBoundedUnder (· ≥ ·) atTop f) : BddBelow (range f) := IsBoundedUnder.bddAbove_range (β := βᵒᵈ) hf -/-- `IsCobounded (≺) f` states that the filter `f` does not tend to infinity w.r.t. `≺`. This is -also called frequently bounded. Will be usually instantiated with `≤` or `≥`. - -There is a subtlety in this definition: we want `f.IsCobounded` to hold for any `f` in the case of -complete lattices. This will be relevant to deduce theorems on complete lattices from their -versions on conditionally complete lattices with additional assumptions. We have to be careful in -the edge case of the trivial filter containing the empty set: the other natural definition - `¬ ∀ a, ∀ᶠ n in f, a ≤ n` -would not work as well in this case. --/ -def IsCobounded (r : α → α → Prop) (f : Filter α) := - ∃ b, ∀ a, (∀ᶠ x in f, r x a) → r b a - -/-- `IsCoboundedUnder (≺) f u` states that the image of the filter `f` under the map `u` does not -tend to infinity w.r.t. `≺`. This is also called frequently bounded. Will be usually instantiated -with `≤` or `≥`. -/ -def IsCoboundedUnder (r : α → α → Prop) (f : Filter β) (u : β → α) := - (map u f).IsCobounded r - /-- To check that a filter is frequently bounded, it suffices to have a witness which bounds `f` at some point for every admissible set. diff --git a/Mathlib/Topology/Bornology/Basic.lean b/Mathlib/Topology/Bornology/Basic.lean index 6b8a44d2bdedf..e8de59a17b369 100644 --- a/Mathlib/Topology/Bornology/Basic.lean +++ b/Mathlib/Topology/Bornology/Basic.lean @@ -261,7 +261,7 @@ open Bornology theorem Filter.HasBasis.disjoint_cobounded_iff [Bornology α] {ι : Sort*} {p : ι → Prop} {s : ι → Set α} {l : Filter α} (h : l.HasBasis p s) : - Disjoint l (cobounded α) ↔ ∃ i, p i ∧ IsBounded (s i) := + Disjoint l (cobounded α) ↔ ∃ i, p i ∧ Bornology.IsBounded (s i) := h.disjoint_iff_left theorem Set.Finite.isBounded [Bornology α] {s : Set α} (hs : s.Finite) : IsBounded s := From ebd640751c72435bd1a7adfbbb1a4ee63c59cd9d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Nov 2024 11:42:19 +0000 Subject: [PATCH 282/829] chore: rename `Codisjoint_comm` to `codisjoint_comm` (#19403) This was a typo --- Mathlib/Data/Fintype/Basic.lean | 2 +- Mathlib/Order/BooleanAlgebra.lean | 2 +- Mathlib/Order/Disjoint.lean | 6 ++++-- Mathlib/Order/Heyting/Basic.lean | 2 +- 4 files changed, 7 insertions(+), 5 deletions(-) diff --git a/Mathlib/Data/Fintype/Basic.lean b/Mathlib/Data/Fintype/Basic.lean index d77529bf9ac3e..d381a6d94e4b9 100644 --- a/Mathlib/Data/Fintype/Basic.lean +++ b/Mathlib/Data/Fintype/Basic.lean @@ -140,7 +140,7 @@ theorem codisjoint_left : Codisjoint s t ↔ ∀ ⦃a⦄, a ∉ s → a ∈ t := classical simp [codisjoint_iff, eq_univ_iff_forall, or_iff_not_imp_left] theorem codisjoint_right : Codisjoint s t ↔ ∀ ⦃a⦄, a ∉ t → a ∈ s := - Codisjoint_comm.trans codisjoint_left + codisjoint_comm.trans codisjoint_left instance booleanAlgebra [DecidableEq α] : BooleanAlgebra (Finset α) := GeneralizedBooleanAlgebra.toBooleanAlgebra diff --git a/Mathlib/Order/BooleanAlgebra.lean b/Mathlib/Order/BooleanAlgebra.lean index a2918657eefb8..cbe13b58400f7 100644 --- a/Mathlib/Order/BooleanAlgebra.lean +++ b/Mathlib/Order/BooleanAlgebra.lean @@ -681,7 +681,7 @@ theorem codisjoint_himp_self_right : Codisjoint x (x ⇨ y) := @disjoint_sdiff_self_right αᵒᵈ _ _ _ theorem himp_le : x ⇨ y ≤ z ↔ y ≤ z ∧ Codisjoint x z := - (@le_sdiff αᵒᵈ _ _ _ _).trans <| and_congr_right' <| @Codisjoint_comm _ (_) _ _ _ + (@le_sdiff αᵒᵈ _ _ _ _).trans <| and_congr_right' <| @codisjoint_comm _ (_) _ _ _ @[simp] lemma himp_le_iff : x ⇨ y ≤ x ↔ x = ⊤ := ⟨fun h ↦ codisjoint_self.1 <| codisjoint_himp_self_right.mono_right h, fun h ↦ le_top.trans h.ge⟩ diff --git a/Mathlib/Order/Disjoint.lean b/Mathlib/Order/Disjoint.lean index b024a335e595c..59bcd0815347d 100644 --- a/Mathlib/Order/Disjoint.lean +++ b/Mathlib/Order/Disjoint.lean @@ -205,12 +205,14 @@ arguments. -/ def Codisjoint (a b : α) : Prop := ∀ ⦃x⦄, a ≤ x → b ≤ x → ⊤ ≤ x -theorem Codisjoint_comm : Codisjoint a b ↔ Codisjoint b a := +theorem codisjoint_comm : Codisjoint a b ↔ Codisjoint b a := forall_congr' fun _ ↦ forall_swap +@[deprecated (since := "2024-11-23")] alias Codisjoint_comm := codisjoint_comm + @[symm] theorem Codisjoint.symm ⦃a b : α⦄ : Codisjoint a b → Codisjoint b a := - Codisjoint_comm.1 + codisjoint_comm.1 theorem symmetric_codisjoint : Symmetric (Codisjoint : α → α → Prop) := Codisjoint.symm diff --git a/Mathlib/Order/Heyting/Basic.lean b/Mathlib/Order/Heyting/Basic.lean index c44de8d3c3d14..1252f8884377c 100644 --- a/Mathlib/Order/Heyting/Basic.lean +++ b/Mathlib/Order/Heyting/Basic.lean @@ -794,7 +794,7 @@ theorem hnot_le_iff_codisjoint_right : ¬a ≤ b ↔ Codisjoint a b := by rw [← top_sdiff', sdiff_le_iff, codisjoint_iff_le_sup] theorem hnot_le_iff_codisjoint_left : ¬a ≤ b ↔ Codisjoint b a := - hnot_le_iff_codisjoint_right.trans Codisjoint_comm + hnot_le_iff_codisjoint_right.trans codisjoint_comm theorem hnot_le_comm : ¬a ≤ b ↔ ¬b ≤ a := by rw [hnot_le_iff_codisjoint_right, hnot_le_iff_codisjoint_left] From f2336c4d027db5bef5813f1a348c8df5513ad5f6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Nov 2024 11:42:20 +0000 Subject: [PATCH 283/829] chore: delete lemmas about T0 seminormed groups (#19438) Those duplicate lemmas about normed groups --- .../Analysis/Distribution/SchwartzSpace.lean | 2 +- .../Analysis/InnerProductSpace/Positive.lean | 2 +- .../Normed/Affine/ContinuousAffineMap.lean | 2 +- Mathlib/Analysis/Normed/Group/Basic.lean | 52 +++++++------------ Mathlib/Analysis/Normed/Group/Uniform.lean | 2 +- .../MeasureTheory/Integral/IntegrableOn.lean | 2 +- .../NumberField/CanonicalEmbedding/Basic.lean | 4 +- scripts/nolints_prime_decls.txt | 5 +- 8 files changed, 31 insertions(+), 40 deletions(-) diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index 0c1edc7031b70..7aab3fe966d3b 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -1172,7 +1172,7 @@ instance instZeroAtInftyContinuousMapClass : ZeroAtInftyContinuousMapClass 𝓢( intro x hx rw [div_lt_iff₀ hε] at hx have hxpos : 0 < ‖x‖ := by - rw [norm_pos_iff'] + rw [norm_pos_iff] intro hxzero simp only [hxzero, norm_zero, zero_mul, ← not_le] at hx exact hx (apply_nonneg (SchwartzMap.seminorm ℝ 1 0) f) diff --git a/Mathlib/Analysis/InnerProductSpace/Positive.lean b/Mathlib/Analysis/InnerProductSpace/Positive.lean index f49a0c4973a85..305ed79703826 100644 --- a/Mathlib/Analysis/InnerProductSpace/Positive.lean +++ b/Mathlib/Analysis/InnerProductSpace/Positive.lean @@ -109,7 +109,7 @@ lemma antilipschitz_of_forall_le_inner_map {H : Type*} [NormedAddCommGroup H] simp_rw [sq, mul_assoc] at h by_cases hx0 : x = 0 · simp [hx0] - · apply (map_le_map_iff <| OrderIso.mulLeft₀ ‖x‖ (norm_pos_iff'.mpr hx0)).mp + · apply (map_le_map_iff <| OrderIso.mulLeft₀ ‖x‖ (norm_pos_iff.mpr hx0)).mp exact (h x).trans <| (norm_inner_le_norm _ _).trans <| (mul_comm _ _).le lemma isUnit_of_forall_le_norm_inner_map (f : E →L[𝕜] E) {c : ℝ≥0} (hc : 0 < c) diff --git a/Mathlib/Analysis/Normed/Affine/ContinuousAffineMap.lean b/Mathlib/Analysis/Normed/Affine/ContinuousAffineMap.lean index 7491177284cf8..b948db9a803a4 100644 --- a/Mathlib/Analysis/Normed/Affine/ContinuousAffineMap.lean +++ b/Mathlib/Analysis/Normed/Affine/ContinuousAffineMap.lean @@ -175,7 +175,7 @@ noncomputable instance : NormedAddCommGroup (V →ᴬ[𝕜] W) := simp only [norm_eq_zero, coe_const, Function.const_apply] at h₁ rw [h₁] rfl - · rw [norm_eq_zero', contLinear_eq_zero_iff_exists_const] at h₁ + · rw [norm_eq_zero, contLinear_eq_zero_iff_exists_const] at h₁ obtain ⟨q, rfl⟩ := h₁ simp only [norm_le_zero_iff, coe_const, Function.const_apply] at h₂ rw [h₂] diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index dd3bf2a5b2797..da9bbaf0e7c48 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -922,20 +922,6 @@ theorem SeminormedCommGroup.mem_closure_iff : a ∈ closure s ↔ ∀ ε, 0 < ε → ∃ b ∈ s, ‖a / b‖ < ε := by simp [Metric.mem_closure_iff, dist_eq_norm_div] -@[to_additive norm_le_zero_iff'] -theorem norm_le_zero_iff''' [T0Space E] {a : E} : ‖a‖ ≤ 0 ↔ a = 1 := by - letI : NormedGroup E := - { ‹SeminormedGroup E› with toMetricSpace := MetricSpace.ofT0PseudoMetricSpace E } - rw [← dist_one_right, dist_le_zero] - -@[to_additive norm_eq_zero'] -theorem norm_eq_zero''' [T0Space E] {a : E} : ‖a‖ = 0 ↔ a = 1 := - (norm_nonneg' a).le_iff_eq.symm.trans norm_le_zero_iff''' - -@[to_additive norm_pos_iff'] -theorem norm_pos_iff''' [T0Space E] {a : E} : 0 < ‖a‖ ↔ a ≠ 1 := by - rw [← not_le, norm_le_zero_iff'''] - @[to_additive] theorem SeminormedGroup.tendstoUniformlyOn_one {f : ι → κ → G} {s : Set κ} {l : Filter ι} : TendstoUniformlyOn f 1 l s ↔ ∀ ε > 0, ∀ᶠ i in l, ∀ x ∈ s, ‖f i x‖ < ε := by @@ -1292,24 +1278,26 @@ section NormedGroup variable [NormedGroup E] {a b : E} +@[to_additive (attr := simp) norm_le_zero_iff] +lemma norm_le_zero_iff' : ‖a‖ ≤ 0 ↔ a = 1 := by rw [← dist_one_right, dist_le_zero] + +@[to_additive (attr := simp) norm_pos_iff] +lemma norm_pos_iff' : 0 < ‖a‖ ↔ a ≠ 1 := by rw [← not_le, norm_le_zero_iff'] + @[to_additive (attr := simp) norm_eq_zero] -theorem norm_eq_zero'' : ‖a‖ = 0 ↔ a = 1 := - norm_eq_zero''' +lemma norm_eq_zero' : ‖a‖ = 0 ↔ a = 1 := (norm_nonneg' a).le_iff_eq.symm.trans norm_le_zero_iff' @[to_additive norm_ne_zero_iff] -theorem norm_ne_zero_iff' : ‖a‖ ≠ 0 ↔ a ≠ 1 := - norm_eq_zero''.not +lemma norm_ne_zero_iff' : ‖a‖ ≠ 0 ↔ a ≠ 1 := norm_eq_zero'.not -@[to_additive (attr := simp) norm_pos_iff] -theorem norm_pos_iff'' : 0 < ‖a‖ ↔ a ≠ 1 := - norm_pos_iff''' - -@[to_additive (attr := simp) norm_le_zero_iff] -theorem norm_le_zero_iff'' : ‖a‖ ≤ 0 ↔ a = 1 := - norm_le_zero_iff''' +@[deprecated (since := "2024-11-24")] alias norm_le_zero_iff'' := norm_le_zero_iff' +@[deprecated (since := "2024-11-24")] alias norm_le_zero_iff''' := norm_le_zero_iff' +@[deprecated (since := "2024-11-24")] alias norm_pos_iff'' := norm_pos_iff' +@[deprecated (since := "2024-11-24")] alias norm_eq_zero'' := norm_eq_zero' +@[deprecated (since := "2024-11-24")] alias norm_eq_zero''' := norm_eq_zero' @[to_additive] -theorem norm_div_eq_zero_iff : ‖a / b‖ = 0 ↔ a = b := by rw [norm_eq_zero'', div_eq_one] +theorem norm_div_eq_zero_iff : ‖a / b‖ = 0 ↔ a = b := by rw [norm_eq_zero', div_eq_one] @[to_additive] theorem norm_div_pos_iff : 0 < ‖a / b‖ ↔ a ≠ b := by @@ -1318,7 +1306,7 @@ theorem norm_div_pos_iff : 0 < ‖a / b‖ ↔ a ≠ b := by @[to_additive eq_of_norm_sub_le_zero] theorem eq_of_norm_div_le_zero (h : ‖a / b‖ ≤ 0) : a = b := by - rwa [← div_eq_one, ← norm_le_zero_iff''] + rwa [← div_eq_one, ← norm_le_zero_iff'] alias ⟨eq_of_norm_div_eq_zero, _⟩ := norm_div_eq_zero_iff @@ -1334,7 +1322,7 @@ theorem eq_one_or_nnnorm_pos (a : E) : a = 1 ∨ 0 < ‖a‖₊ := @[to_additive (attr := simp) nnnorm_eq_zero] theorem nnnorm_eq_zero' : ‖a‖₊ = 0 ↔ a = 1 := by - rw [← NNReal.coe_eq_zero, coe_nnnorm', norm_eq_zero''] + rw [← NNReal.coe_eq_zero, coe_nnnorm', norm_eq_zero'] @[to_additive nnnorm_ne_zero_iff] theorem nnnorm_ne_zero_iff' : ‖a‖₊ ≠ 0 ↔ a ≠ 1 := @@ -1346,24 +1334,24 @@ lemma nnnorm_pos' : 0 < ‖a‖₊ ↔ a ≠ 1 := pos_iff_ne_zero.trans nnnorm_n /-- See `tendsto_norm_one` for a version with full neighborhoods. -/ @[to_additive "See `tendsto_norm_zero` for a version with full neighborhoods."] lemma tendsto_norm_one' : Tendsto (norm : E → ℝ) (𝓝[≠] 1) (𝓝[>] 0) := - tendsto_norm_one.inf <| tendsto_principal_principal.2 fun _ hx ↦ norm_pos_iff''.2 hx + tendsto_norm_one.inf <| tendsto_principal_principal.2 fun _ hx ↦ norm_pos_iff'.2 hx @[to_additive] theorem tendsto_norm_div_self_punctured_nhds (a : E) : Tendsto (fun x => ‖x / a‖) (𝓝[≠] a) (𝓝[>] 0) := (tendsto_norm_div_self a).inf <| - tendsto_principal_principal.2 fun _x hx => norm_pos_iff''.2 <| div_ne_one.2 hx + tendsto_principal_principal.2 fun _x hx => norm_pos_iff'.2 <| div_ne_one.2 hx @[to_additive] theorem tendsto_norm_nhdsWithin_one : Tendsto (norm : E → ℝ) (𝓝[≠] 1) (𝓝[>] 0) := - tendsto_norm_one.inf <| tendsto_principal_principal.2 fun _x => norm_pos_iff''.2 + tendsto_norm_one.inf <| tendsto_principal_principal.2 fun _x => norm_pos_iff'.2 variable (E) /-- The norm of a normed group as a group norm. -/ @[to_additive "The norm of a normed group as an additive group norm."] def normGroupNorm : GroupNorm E := - { normGroupSeminorm _ with eq_one_of_map_eq_zero' := fun _ => norm_eq_zero''.1 } + { normGroupSeminorm _ with eq_one_of_map_eq_zero' := fun _ => norm_eq_zero'.1 } @[simp] theorem coe_normGroupNorm : ⇑(normGroupNorm E) = norm := diff --git a/Mathlib/Analysis/Normed/Group/Uniform.lean b/Mathlib/Analysis/Normed/Group/Uniform.lean index 73b33458e9210..ca6dd25e1c9bb 100644 --- a/Mathlib/Analysis/Normed/Group/Uniform.lean +++ b/Mathlib/Analysis/Normed/Group/Uniform.lean @@ -381,7 +381,7 @@ instance : NormedCommGroup (SeparationQuotient E) where @[to_additive] theorem mk_eq_one_iff {p : E} : mk p = 1 ↔ ‖p‖ = 0 := by - rw [← norm_mk', norm_eq_zero''] + rw [← norm_mk', norm_eq_zero'] set_option linter.docPrime false in @[to_additive (attr := simp) nnnorm_mk] diff --git a/Mathlib/MeasureTheory/Integral/IntegrableOn.lean b/Mathlib/MeasureTheory/Integral/IntegrableOn.lean index f9d75e6cfea42..d23df43813b60 100644 --- a/Mathlib/MeasureTheory/Integral/IntegrableOn.lean +++ b/Mathlib/MeasureTheory/Integral/IntegrableOn.lean @@ -489,7 +489,7 @@ lemma IntegrableAtFilter.eq_zero_of_tendsto (h : IntegrableAtFilter f l μ) (h' : ∀ s ∈ l, μ s = ∞) {a : E} (hf : Tendsto f l (𝓝 a)) : a = 0 := by by_contra H - obtain ⟨ε, εpos, hε⟩ : ∃ (ε : ℝ), 0 < ε ∧ ε < ‖a‖ := exists_between (norm_pos_iff'.mpr H) + obtain ⟨ε, εpos, hε⟩ : ∃ (ε : ℝ), 0 < ε ∧ ε < ‖a‖ := exists_between (norm_pos_iff.mpr H) rcases h with ⟨u, ul, hu⟩ let v := u ∩ {b | ε < ‖f b‖} have hv : IntegrableOn f v μ := hu.mono_set inter_subset_left diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean index 3cfe928f41764..84f4e2aa14e91 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean @@ -367,8 +367,8 @@ theorem forall_normAtPlace_eq_zero_iff {x : mixedSpace K} : (∀ w, normAtPlace w x = 0) ↔ x = 0 := by refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ · ext w - · exact norm_eq_zero'.mp (normAtPlace_apply_isReal w.prop _ ▸ h w.1) - · exact norm_eq_zero'.mp (normAtPlace_apply_isComplex w.prop _ ▸ h w.1) + · exact norm_eq_zero.mp (normAtPlace_apply_isReal w.prop _ ▸ h w.1) + · exact norm_eq_zero.mp (normAtPlace_apply_isComplex w.prop _ ▸ h w.1) · simp_rw [h, map_zero, implies_true] @[deprecated (since := "2024-09-13")] alias normAtPlace_eq_zero := forall_normAtPlace_eq_zero_iff diff --git a/scripts/nolints_prime_decls.txt b/scripts/nolints_prime_decls.txt index a5ca83037daf6..414151f344621 100644 --- a/scripts/nolints_prime_decls.txt +++ b/scripts/nolints_prime_decls.txt @@ -3455,6 +3455,7 @@ NormedSpace.isVonNBounded_iff' NormedSpace.norm_expSeries_summable' NormedSpace.norm_expSeries_summable_of_mem_ball' norm_eq_of_mem_sphere' +norm_eq_zero' norm_eq_zero'' norm_eq_zero''' norm_inv' @@ -3462,6 +3463,7 @@ norm_le_norm_add_const_of_dist_le' norm_le_norm_add_norm_div' norm_le_of_mem_closedBall' norm_le_pi_norm' +norm_le_zero_iff' norm_le_zero_iff'' norm_le_zero_iff''' norm_lt_of_mem_ball' @@ -3469,7 +3471,8 @@ norm_ne_zero_iff' norm_nonneg' norm_of_subsingleton' norm_one' -norm_pos_iff'' +norm_pos_iff' +norm_pos_iff' norm_pos_iff''' norm_sub_norm_le' norm_toNNReal' From 85ded51450df1dca83cdb8736849a0028a226dcc Mon Sep 17 00:00:00 2001 From: sgouezel Date: Mon, 25 Nov 2024 11:42:22 +0000 Subject: [PATCH 284/829] =?UTF-8?q?feat:=20extend=20`ContDiff`=20to=20take?= =?UTF-8?q?=20a=20smoothness=20exponent=20in=20`WithTop=20=E2=84=95?= =?UTF-8?q?=E2=88=9E`=20(#19442)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This is mathematically not meaningful at all, since the definition of `C^∞` and `C^ω` functions is the same in this PR. However, this opens the way to replacing the definition of `C^ω` with analytic in #17152. So, the current PR is just an intermediate step towards #17152, but making sure that the final diff will be smaller than the current +1280 lines changes. There is *no* mathematical content in the current PR: it's mostly adapting the library to provide an element of `WithTop ℕ∞` when it was expecting an element of `ℕ∞` before. In the locale `ContDiff`, we introduce two notations `ω` for `⊤ : WithTop ℕ∞ ` and `∞` for `(⊤ : ℕ∞) : WithTop ℕ∞`. A lot of the changes are to open the locale and use `∞` at suitable places. --- Archive/Hairer.lean | 3 +- .../Calculus/AddTorsor/AffineMap.lean | 2 +- .../Analysis/Calculus/BumpFunction/Basic.lean | 7 +- .../BumpFunction/FiniteDimension.lean | 27 +-- Mathlib/Analysis/Calculus/ContDiff/Basic.lean | 170 ++++++++++-------- .../Analysis/Calculus/ContDiff/Bounds.lean | 77 ++++---- Mathlib/Analysis/Calculus/ContDiff/Defs.lean | 114 +++++++----- .../Calculus/ContDiff/FiniteDimension.lean | 7 +- .../Analysis/Calculus/ContDiff/RCLike.lean | 8 +- .../Analysis/Calculus/FDeriv/Analytic.lean | 34 ++-- Mathlib/Analysis/Calculus/FDeriv/Norm.lean | 12 +- .../Analysis/Calculus/FDeriv/Symmetric.lean | 4 +- .../InverseFunctionTheorem/ContDiff.lean | 16 +- .../Analysis/Calculus/IteratedDeriv/Defs.lean | 22 +-- Mathlib/Analysis/Calculus/Rademacher.lean | 7 +- Mathlib/Analysis/Calculus/SmoothSeries.lean | 25 ++- Mathlib/Analysis/Calculus/Taylor.lean | 7 +- Mathlib/Analysis/Calculus/VectorField.lean | 8 +- Mathlib/Analysis/Complex/RealDeriv.lean | 4 +- Mathlib/Analysis/Convolution.lean | 9 +- .../Distribution/AEEqOfIntegralContDiff.lean | 8 +- .../Analysis/Distribution/SchwartzSpace.lean | 45 +++-- .../Fourier/FourierTransformDeriv.lean | 16 +- .../Analysis/SpecialFunctions/ExpDeriv.lean | 4 +- .../Analysis/SpecialFunctions/Log/Deriv.lean | 9 +- .../Analysis/SpecialFunctions/Pow/Deriv.lean | 7 +- .../SpecialFunctions/SmoothTransition.lean | 6 +- .../Trigonometric/ComplexDeriv.lean | 2 +- .../Trigonometric/InverseDeriv.lean | 21 ++- .../Algebra/LeftInvariantDerivation.lean | 3 + Mathlib/Geometry/Manifold/Algebra/Monoid.lean | 3 + .../Manifold/Algebra/SmoothFunctions.lean | 3 + .../Geometry/Manifold/Algebra/Structures.lean | 4 +- .../Geometry/Manifold/ContMDiff/Atlas.lean | 22 +-- .../Geometry/Manifold/ContMDiff/Basic.lean | 2 +- Mathlib/Geometry/Manifold/ContMDiff/Defs.lean | 9 +- .../Geometry/Manifold/ContMDiffMFDeriv.lean | 5 +- Mathlib/Geometry/Manifold/ContMDiffMap.lean | 4 + .../Geometry/Manifold/DerivationBundle.lean | 3 + Mathlib/Geometry/Manifold/Diffeomorph.lean | 4 +- Mathlib/Geometry/Manifold/Instances/Real.lean | 2 +- .../Geometry/Manifold/Instances/Sphere.lean | 26 +-- .../Instances/UnitsOfNormedAlgebra.lean | 3 + Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean | 6 +- Mathlib/Geometry/Manifold/MFDeriv/Basic.lean | 3 +- Mathlib/Geometry/Manifold/MFDeriv/Defs.lean | 9 +- .../Geometry/Manifold/PartitionOfUnity.lean | 3 + .../Manifold/Sheaf/LocallyRingedSpace.lean | 4 + Mathlib/Geometry/Manifold/Sheaf/Smooth.lean | 4 + .../Manifold/SmoothManifoldWithCorners.lean | 34 ++-- .../Manifold/VectorBundle/SmoothSection.lean | 3 + .../Manifold/VectorBundle/Tangent.lean | 11 +- .../Geometry/Manifold/WhitneyEmbedding.lean | 3 + MathlibTest/fun_prop2.lean | 2 +- 54 files changed, 483 insertions(+), 373 deletions(-) diff --git a/Archive/Hairer.lean b/Archive/Hairer.lean index e9b91c32132a8..6abb80f254220 100644 --- a/Archive/Hairer.lean +++ b/Archive/Hairer.lean @@ -26,6 +26,7 @@ noncomputable section open Metric Set MeasureTheory open MvPolynomial hiding support open Function hiding eval +open scoped ContDiff variable {ι : Type*} [Fintype ι] @@ -110,7 +111,7 @@ lemma inj_L : Injective (L ι) := apply subset_closure lemma hairer (N : ℕ) (ι : Type*) [Fintype ι] : - ∃ (ρ : EuclideanSpace ℝ ι → ℝ), tsupport ρ ⊆ closedBall 0 1 ∧ ContDiff ℝ ⊤ ρ ∧ + ∃ (ρ : EuclideanSpace ℝ ι → ℝ), tsupport ρ ⊆ closedBall 0 1 ∧ ContDiff ℝ ∞ ρ ∧ ∀ (p : MvPolynomial ι ℝ), p.totalDegree ≤ N → ∫ x : EuclideanSpace ℝ ι, eval x p • ρ x = eval 0 p := by have := (inj_L ι).comp (restrictTotalDegree ι ℝ N).injective_subtype diff --git a/Mathlib/Analysis/Calculus/AddTorsor/AffineMap.lean b/Mathlib/Analysis/Calculus/AddTorsor/AffineMap.lean index 5eb7534bb2ff4..104a452fe5423 100644 --- a/Mathlib/Analysis/Calculus/AddTorsor/AffineMap.lean +++ b/Mathlib/Analysis/Calculus/AddTorsor/AffineMap.lean @@ -25,7 +25,7 @@ variable [NormedAddCommGroup V] [NormedSpace 𝕜 V] variable [NormedAddCommGroup W] [NormedSpace 𝕜 W] /-- A continuous affine map between normed vector spaces is smooth. -/ -theorem contDiff {n : ℕ∞} (f : V →ᴬ[𝕜] W) : ContDiff 𝕜 n f := by +theorem contDiff {n : WithTop ℕ∞} (f : V →ᴬ[𝕜] W) : ContDiff 𝕜 n f := by rw [f.decomp] apply f.contLinear.contDiff.add exact contDiff_const diff --git a/Mathlib/Analysis/Calculus/BumpFunction/Basic.lean b/Mathlib/Analysis/Calculus/BumpFunction/Basic.lean index 7c20ab571383f..922effc6f97d8 100644 --- a/Mathlib/Analysis/Calculus/BumpFunction/Basic.lean +++ b/Mathlib/Analysis/Calculus/BumpFunction/Basic.lean @@ -49,7 +49,7 @@ smooth function, smooth bump function noncomputable section open Function Set Filter -open scoped Topology Filter +open scoped Topology Filter ContDiff variable {E X : Type*} @@ -80,7 +80,7 @@ structure ContDiffBumpBase (E : Type*) [NormedAddCommGroup E] [NormedSpace ℝ E toFun : ℝ → E → ℝ mem_Icc : ∀ (R : ℝ) (x : E), toFun R x ∈ Icc (0 : ℝ) 1 symmetric : ∀ (R : ℝ) (x : E), toFun R (-x) = toFun R x - smooth : ContDiffOn ℝ ⊤ (uncurry toFun) (Ioi (1 : ℝ) ×ˢ (univ : Set E)) + smooth : ContDiffOn ℝ ∞ (uncurry toFun) (Ioi (1 : ℝ) ×ˢ (univ : Set E)) eq_one : ∀ R : ℝ, 1 < R → ∀ x : E, ‖x‖ ≤ 1 → toFun R x = 1 support : ∀ R : ℝ, 1 < R → Function.support (toFun R) = Metric.ball (0 : E) R @@ -178,7 +178,8 @@ protected theorem _root_.ContDiffWithinAt.contDiffBump {c g : X → E} {s : Set ContDiffWithinAt ℝ n (fun x => f x (g x)) s x := by change ContDiffWithinAt ℝ n (uncurry (someContDiffBumpBase E).toFun ∘ fun x : X => ((f x).rOut / (f x).rIn, (f x).rIn⁻¹ • (g x - c x))) s x - refine (((someContDiffBumpBase E).smooth.contDiffAt ?_).of_le le_top).comp_contDiffWithinAt x ?_ + refine (((someContDiffBumpBase E).smooth.contDiffAt ?_).of_le + (mod_cast le_top)).comp_contDiffWithinAt x ?_ · exact prod_mem_nhds (Ioi_mem_nhds (f x).one_lt_rOut_div_rIn) univ_mem · exact (hR.div hr (f x).rIn_pos.ne').prod ((hr.inv (f x).rIn_pos.ne').smul (hg.sub hc)) diff --git a/Mathlib/Analysis/Calculus/BumpFunction/FiniteDimension.lean b/Mathlib/Analysis/Calculus/BumpFunction/FiniteDimension.lean index d0f390973978c..2f3e0be74bdfb 100644 --- a/Mathlib/Analysis/Calculus/BumpFunction/FiniteDimension.lean +++ b/Mathlib/Analysis/Calculus/BumpFunction/FiniteDimension.lean @@ -28,7 +28,7 @@ noncomputable section open Set Metric TopologicalSpace Function Asymptotics MeasureTheory Module ContinuousLinearMap Filter MeasureTheory.Measure Bornology -open scoped Pointwise Topology NNReal Convolution +open scoped Pointwise Topology NNReal Convolution ContDiff variable {E : Type*} [NormedAddCommGroup E] @@ -40,7 +40,7 @@ variable [NormedSpace ℝ E] [FiniteDimensional ℝ E] values in `[0, 1]`, supported in `s` and with `f x = 1`. -/ theorem exists_smooth_tsupport_subset {s : Set E} {x : E} (hs : s ∈ 𝓝 x) : ∃ f : E → ℝ, - tsupport f ⊆ s ∧ HasCompactSupport f ∧ ContDiff ℝ ⊤ f ∧ range f ⊆ Icc 0 1 ∧ f x = 1 := by + tsupport f ⊆ s ∧ HasCompactSupport f ∧ ContDiff ℝ ∞ f ∧ range f ⊆ Icc 0 1 ∧ f x = 1 := by obtain ⟨d : ℝ, d_pos : 0 < d, hd : Euclidean.closedBall x d ⊆ s⟩ := Euclidean.nhds_basis_closedBall.mem_iff.1 hs let c : ContDiffBump (toEuclidean x) := @@ -73,7 +73,7 @@ theorem exists_smooth_tsupport_subset {s : Set E} {x : E} (hs : s ∈ 𝓝 x) : /-- Given an open set `s` in a finite-dimensional real normed vector space, there exists a smooth function with values in `[0, 1]` whose support is exactly `s`. -/ theorem IsOpen.exists_smooth_support_eq {s : Set E} (hs : IsOpen s) : - ∃ f : E → ℝ, f.support = s ∧ ContDiff ℝ ⊤ f ∧ Set.range f ⊆ Set.Icc 0 1 := by + ∃ f : E → ℝ, f.support = s ∧ ContDiff ℝ ∞ f ∧ Set.range f ⊆ Set.Icc 0 1 := by /- For any given point `x` in `s`, one can construct a smooth function with support in `s` and nonzero at `x`. By second-countability, it follows that we may cover `s` with the supports of countably many such functions, say `g i`. @@ -85,7 +85,7 @@ theorem IsOpen.exists_smooth_support_eq {s : Set E} (hs : IsOpen s) : · exact ⟨fun _ => 0, Function.support_zero, contDiff_const, by simp only [range_const, singleton_subset_iff, left_mem_Icc, zero_le_one]⟩ - let ι := { f : E → ℝ // f.support ⊆ s ∧ HasCompactSupport f ∧ ContDiff ℝ ⊤ f ∧ range f ⊆ Icc 0 1 } + let ι := { f : E → ℝ // f.support ⊆ s ∧ HasCompactSupport f ∧ ContDiff ℝ ∞ f ∧ range f ⊆ Icc 0 1 } obtain ⟨T, T_count, hT⟩ : ∃ T : Set ι, T.Countable ∧ ⋃ f ∈ T, support (f : E → ℝ) = s := by have : ⋃ f : ι, (f : E → ℝ).support = s := by refine Subset.antisymm (iUnion_subset fun f => f.2.1) ?_ @@ -116,7 +116,7 @@ theorem IsOpen.exists_smooth_support_eq {s : Set E} (hs : IsOpen s) : rcases iT with ⟨n, hn⟩ rw [← hn] at hi exact ⟨n, hi⟩ - have g_smooth : ∀ n, ContDiff ℝ ⊤ (g n) := fun n => (g0 n).2.2.2.1 + have g_smooth : ∀ n, ContDiff ℝ ∞ (g n) := fun n => (g0 n).2.2.2.1 have g_comp_supp : ∀ n, HasCompactSupport (g n) := fun n => (g0 n).2.2.1 have g_nonneg : ∀ n x, 0 ≤ g n x := fun n x => ((g0 n).2.2.2.2 (mem_range_self x)).1 obtain ⟨δ, δpos, c, δc, c_lt⟩ : @@ -127,8 +127,8 @@ theorem IsOpen.exists_smooth_support_eq {s : Set E} (hs : IsOpen s) : have : ∀ i, ∃ R, ∀ x, ‖iteratedFDeriv ℝ i (fun x => g n x) x‖ ≤ R := by intro i have : BddAbove (range fun x => ‖iteratedFDeriv ℝ i (fun x : E => g n x) x‖) := by - apply - ((g_smooth n).continuous_iteratedFDeriv le_top).norm.bddAbove_range_of_hasCompactSupport + apply ((g_smooth n).continuous_iteratedFDeriv + (mod_cast le_top)).norm.bddAbove_range_of_hasCompactSupport apply HasCompactSupport.comp_left _ norm_zero apply (g_comp_supp n).iteratedFDeriv rcases this with ⟨R, hR⟩ @@ -148,7 +148,8 @@ theorem IsOpen.exists_smooth_support_eq {s : Set E} (hs : IsOpen s) : refine ⟨M⁻¹ * δ n, by positivity, fun i hi x => ?_⟩ calc ‖iteratedFDeriv ℝ i ((M⁻¹ * δ n) • g n) x‖ = ‖(M⁻¹ * δ n) • iteratedFDeriv ℝ i (g n) x‖ := by - rw [iteratedFDeriv_const_smul_apply]; exact (g_smooth n).of_le le_top + rw [iteratedFDeriv_const_smul_apply] + exact (g_smooth n).of_le (mod_cast le_top) _ = M⁻¹ * δ n * ‖iteratedFDeriv ℝ i (g n) x‖ := by rw [norm_smul _ (iteratedFDeriv ℝ i (g n) x), Real.norm_of_nonneg]; positivity _ ≤ M⁻¹ * δ n * M := (mul_le_mul_of_nonneg_left ((hR i x).trans (IR i hi)) (by positivity)) @@ -207,10 +208,10 @@ variable (E) theorem u_exists : ∃ u : E → ℝ, - ContDiff ℝ ⊤ u ∧ (∀ x, u x ∈ Icc (0 : ℝ) 1) ∧ support u = ball 0 1 ∧ ∀ x, u (-x) = u x := by + ContDiff ℝ ∞ u ∧ (∀ x, u x ∈ Icc (0 : ℝ) 1) ∧ support u = ball 0 1 ∧ ∀ x, u (-x) = u x := by have A : IsOpen (ball (0 : E) 1) := isOpen_ball obtain ⟨f, f_support, f_smooth, f_range⟩ : - ∃ f : E → ℝ, f.support = ball (0 : E) 1 ∧ ContDiff ℝ ⊤ f ∧ Set.range f ⊆ Set.Icc 0 1 := + ∃ f : E → ℝ, f.support = ball (0 : E) 1 ∧ ContDiff ℝ ∞ f ∧ Set.range f ⊆ Set.Icc 0 1 := A.exists_smooth_support_eq have B : ∀ x, f x ∈ Icc (0 : ℝ) 1 := fun x => f_range (mem_range_self x) refine ⟨fun x => (f x + f (-x)) / 2, ?_, ?_, ?_, ?_⟩ @@ -243,7 +244,7 @@ def u (x : E) : ℝ := variable (E) -theorem u_smooth : ContDiff ℝ ⊤ (u : E → ℝ) := +theorem u_smooth : ContDiff ℝ ∞ (u : E → ℝ) := (Classical.choose_spec (u_exists E)).1 theorem u_continuous : Continuous (u : E → ℝ) := @@ -433,7 +434,7 @@ theorem y_pos_of_mem_ball {D : ℝ} {x : E} (Dpos : 0 < D) (D_lt_one : D < 1) variable (E) -theorem y_smooth : ContDiffOn ℝ ⊤ (uncurry y) (Ioo (0 : ℝ) 1 ×ˢ (univ : Set E)) := by +theorem y_smooth : ContDiffOn ℝ ∞ (uncurry y) (Ioo (0 : ℝ) 1 ×ˢ (univ : Set E)) := by have hs : IsOpen (Ioo (0 : ℝ) (1 : ℝ)) := isOpen_Ioo have hk : IsCompact (closedBall (0 : E) 1) := ProperSpace.isCompact_closedBall _ _ refine contDiffOn_convolution_left_with_param (lsmul ℝ ℝ) hs hk ?_ ?_ ?_ @@ -489,7 +490,7 @@ instance (priority := 100) {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E · rfl smooth := by suffices - ContDiffOn ℝ ⊤ + ContDiffOn ℝ ∞ (uncurry y ∘ fun p : ℝ × E => ((p.1 - 1) / (p.1 + 1), ((p.1 + 1) / 2)⁻¹ • p.2)) (Ioi 1 ×ˢ univ) by apply this.congr diff --git a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean index 52ac217ea8cf2..ad3ed57edf02c 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean @@ -26,7 +26,7 @@ Similar results are given for `C^n` functions on domains. We use the notation `E [×n]→L[𝕜] F` for the space of continuous multilinear maps on `E^n` with values in `F`. This is the space in which the `n`-th derivative of a function from `E` to `F` lives. -In this file, we denote `⊤ : ℕ∞` with `∞`. +In this file, we denote `(⊤ : ℕ∞) : WithTop ℕ∞` with `∞` and `⊤ : WithTop ℕ∞` with `ω`. ## Tags @@ -35,9 +35,7 @@ derivative, differentiability, higher derivative, `C^n`, multilinear, Taylor ser noncomputable section -open scoped NNReal Nat - -local notation "∞" => (⊤ : ℕ∞) +open scoped NNReal Nat ContDiff universe u uE uF uG @@ -52,7 +50,7 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type uE} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {F : Type uF} [NormedAddCommGroup F] [NormedSpace 𝕜 F] {G : Type uG} [NormedAddCommGroup G] [NormedSpace 𝕜 G] {X : Type*} [NormedAddCommGroup X] [NormedSpace 𝕜 X] {s t : Set E} {f : E → F} - {g : F → G} {x x₀ : E} {b : E × F → G} {m n : ℕ∞} {p : E → FormalMultilinearSeries 𝕜 E F} + {g : F → G} {x x₀ : E} {b : E × F → G} {m n : WithTop ℕ∞} {p : E → FormalMultilinearSeries 𝕜 E F} /-! ### Constants -/ @@ -72,16 +70,18 @@ theorem iteratedFDeriv_zero_fun {n : ℕ} : (iteratedFDeriv 𝕜 n fun _ : E ↦ funext fun x ↦ by simpa [← iteratedFDerivWithin_univ] using iteratedFDerivWithin_zero_fun uniqueDiffOn_univ (mem_univ x) -theorem contDiff_zero_fun : ContDiff 𝕜 n fun _ : E => (0 : F) := - contDiff_of_differentiable_iteratedFDeriv fun m _ => by - rw [iteratedFDeriv_zero_fun] - exact differentiable_const (0 : E[×m]→L[𝕜] F) +theorem contDiff_zero_fun : ContDiff 𝕜 n fun _ : E => (0 : F) := by + suffices ContDiff 𝕜 ω (fun _ : E => (0 : F)) from this.of_le le_top + rw [← contDiff_infty_iff_contDiff_omega] + apply contDiff_of_differentiable_iteratedFDeriv fun m _ ↦ ?_ + rw [iteratedFDeriv_zero_fun] + exact differentiable_const (0 : E[×m]→L[𝕜] F) /-- Constants are `C^∞`. -/ theorem contDiff_const {c : F} : ContDiff 𝕜 n fun _ : E => c := by - suffices h : ContDiff 𝕜 ∞ fun _ : E => c from h.of_le le_top - rw [contDiff_top_iff_fderiv] + suffices h : ContDiff 𝕜 ω fun _ : E => c from h.of_le le_top + rw [← contDiff_infty_iff_contDiff_omega, contDiff_top_iff_fderiv] refine ⟨differentiable_const c, ?_⟩ rw [fderiv_const] exact contDiff_zero_fun @@ -144,8 +144,8 @@ theorem contDiffWithinAt_singleton : ContDiffWithinAt 𝕜 n f {x} x := /-- Unbundled bounded linear functions are `C^∞`. -/ theorem IsBoundedLinearMap.contDiff (hf : IsBoundedLinearMap 𝕜 f) : ContDiff 𝕜 n f := by - suffices h : ContDiff 𝕜 ∞ f from h.of_le le_top - rw [contDiff_top_iff_fderiv] + suffices h : ContDiff 𝕜 ω f from h.of_le le_top + rw [← contDiff_infty_iff_contDiff_omega, contDiff_top_iff_fderiv] refine ⟨hf.differentiable, ?_⟩ simp_rw [hf.fderiv] exact contDiff_const @@ -179,8 +179,8 @@ theorem contDiffOn_id {s} : ContDiffOn 𝕜 n (id : E → E) s := /-- Bilinear functions are `C^∞`. -/ theorem IsBoundedBilinearMap.contDiff (hb : IsBoundedBilinearMap 𝕜 b) : ContDiff 𝕜 n b := by - suffices h : ContDiff 𝕜 ∞ b from h.of_le le_top - rw [contDiff_top_iff_fderiv] + suffices h : ContDiff 𝕜 ω b from h.of_le le_top + rw [← contDiff_infty_iff_contDiff_omega, contDiff_top_iff_fderiv] refine ⟨hb.differentiable, ?_⟩ simp only [hb.fderiv] exact hb.isBoundedLinearMap_deriv.contDiff @@ -883,8 +883,8 @@ section ClmApplyConst /-- Application of a `ContinuousLinearMap` to a constant commutes with `iteratedFDerivWithin`. -/ theorem iteratedFDerivWithin_clm_apply_const_apply - {s : Set E} (hs : UniqueDiffOn 𝕜 s) {n : ℕ∞} {c : E → F →L[𝕜] G} (hc : ContDiffOn 𝕜 n c s) - {i : ℕ} (hi : i ≤ n) {x : E} (hx : x ∈ s) {u : F} {m : Fin i → E} : + {s : Set E} (hs : UniqueDiffOn 𝕜 s) {c : E → F →L[𝕜] G} + (hc : ContDiffOn 𝕜 n c s) {i : ℕ} (hi : i ≤ n) {x : E} (hx : x ∈ s) {u : F} {m : Fin i → E} : (iteratedFDerivWithin 𝕜 i (fun y ↦ (c y) u) s x) m = (iteratedFDerivWithin 𝕜 i c s x) m u := by induction i generalizing x with | zero => simp @@ -905,7 +905,7 @@ theorem iteratedFDerivWithin_clm_apply_const_apply /-- Application of a `ContinuousLinearMap` to a constant commutes with `iteratedFDeriv`. -/ theorem iteratedFDeriv_clm_apply_const_apply - {n : ℕ∞} {c : E → F →L[𝕜] G} (hc : ContDiff 𝕜 n c) + {c : E → F →L[𝕜] G} (hc : ContDiff 𝕜 n c) {i : ℕ} (hi : i ≤ n) {x : E} {u : F} {m : Fin i → E} : (iteratedFDeriv 𝕜 i (fun y ↦ (c y) u) x) m = (iteratedFDeriv 𝕜 i c x) m u := by simp only [← iteratedFDerivWithin_univ] @@ -977,7 +977,7 @@ To show that `x ↦ D_yf(x,y)g(x)` (taken within `t`) is `C^m` at `x₀` within * `g` is `C^m` at `x₀` within `s`; * Derivatives are unique at `g(x)` within `t` for `x` sufficiently close to `x₀` within `s ∪ {x₀}`; * `t` is a neighborhood of `g(x₀)` within `g '' s`; -/ -theorem ContDiffWithinAt.fderivWithin'' {f : E → F → G} {g : E → F} {t : Set F} {n : ℕ∞} +theorem ContDiffWithinAt.fderivWithin'' {f : E → F → G} {g : E → F} {t : Set F} (hf : ContDiffWithinAt 𝕜 n (Function.uncurry f) (insert x₀ s ×ˢ t) (x₀, g x₀)) (hg : ContDiffWithinAt 𝕜 m g s x₀) (ht : ∀ᶠ x in 𝓝[insert x₀ s] x₀, UniqueDiffWithinAt 𝕜 t (g x)) (hmn : m + 1 ≤ n) @@ -990,14 +990,18 @@ theorem ContDiffWithinAt.fderivWithin'' {f : E → F → G} {g : E → F} {t : S refine hf'.congr_of_eventuallyEq_insert ?_ filter_upwards [hv, ht] exact fun y hy h2y => (hvf' y hy).fderivWithin h2y - induction' m with m - · obtain rfl := eq_top_iff.mpr hmn + match m with + | ω => + intro k hk + apply this k hk + exact le_rfl + | ∞ => rw [contDiffWithinAt_top] - exact fun m => this m le_top - exact this _ le_rfl + exact fun m => this m (mod_cast le_top) + | (m : ℕ) => exact this _ le_rfl /-- A special case of `ContDiffWithinAt.fderivWithin''` where we require that `s ⊆ g⁻¹(t)`. -/ -theorem ContDiffWithinAt.fderivWithin' {f : E → F → G} {g : E → F} {t : Set F} {n : ℕ∞} +theorem ContDiffWithinAt.fderivWithin' {f : E → F → G} {g : E → F} {t : Set F} (hf : ContDiffWithinAt 𝕜 n (Function.uncurry f) (insert x₀ s ×ˢ t) (x₀, g x₀)) (hg : ContDiffWithinAt 𝕜 m g s x₀) (ht : ∀ᶠ x in 𝓝[insert x₀ s] x₀, UniqueDiffWithinAt 𝕜 t (g x)) (hmn : m + 1 ≤ n) @@ -1006,7 +1010,7 @@ theorem ContDiffWithinAt.fderivWithin' {f : E → F → G} {g : E → F} {t : Se /-- A special case of `ContDiffWithinAt.fderivWithin'` where we require that `x₀ ∈ s` and there are unique derivatives everywhere within `t`. -/ -protected theorem ContDiffWithinAt.fderivWithin {f : E → F → G} {g : E → F} {t : Set F} {n : ℕ∞} +protected theorem ContDiffWithinAt.fderivWithin {f : E → F → G} {g : E → F} {t : Set F} (hf : ContDiffWithinAt 𝕜 n (Function.uncurry f) (s ×ˢ t) (x₀, g x₀)) (hg : ContDiffWithinAt 𝕜 m g s x₀) (ht : UniqueDiffOn 𝕜 t) (hmn : m + 1 ≤ n) (hx₀ : x₀ ∈ s) (hst : s ⊆ g ⁻¹' t) : ContDiffWithinAt 𝕜 m (fun x => fderivWithin 𝕜 (f x) t (g x)) s x₀ := by @@ -1016,7 +1020,7 @@ protected theorem ContDiffWithinAt.fderivWithin {f : E → F → G} {g : E → F exact eventually_of_mem self_mem_nhdsWithin fun x hx => ht _ (hst hx) /-- `x ↦ fderivWithin 𝕜 (f x) t (g x) (k x)` is smooth at a point within a set. -/ -theorem ContDiffWithinAt.fderivWithin_apply {f : E → F → G} {g k : E → F} {t : Set F} {n : ℕ∞} +theorem ContDiffWithinAt.fderivWithin_apply {f : E → F → G} {g k : E → F} {t : Set F} (hf : ContDiffWithinAt 𝕜 n (Function.uncurry f) (s ×ˢ t) (x₀, g x₀)) (hg : ContDiffWithinAt 𝕜 m g s x₀) (hk : ContDiffWithinAt 𝕜 m k s x₀) (ht : UniqueDiffOn 𝕜 t) (hmn : m + 1 ≤ n) (hx₀ : x₀ ∈ s) (hst : s ⊆ g ⁻¹' t) : @@ -1026,7 +1030,7 @@ theorem ContDiffWithinAt.fderivWithin_apply {f : E → F → G} {g k : E → F} /-- `fderivWithin 𝕜 f s` is smooth at `x₀` within `s`. -/ theorem ContDiffWithinAt.fderivWithin_right (hf : ContDiffWithinAt 𝕜 n f s x₀) - (hs : UniqueDiffOn 𝕜 s) (hmn : (m + 1 : ℕ∞) ≤ n) (hx₀s : x₀ ∈ s) : + (hs : UniqueDiffOn 𝕜 s) (hmn : m + 1 ≤ n) (hx₀s : x₀ ∈ s) : ContDiffWithinAt 𝕜 m (fderivWithin 𝕜 f s) s x₀ := ContDiffWithinAt.fderivWithin (ContDiffWithinAt.comp (x₀, x₀) hf contDiffWithinAt_snd <| prod_subset_preimage_snd s s) @@ -1036,7 +1040,7 @@ theorem ContDiffWithinAt.fderivWithin_right (hf : ContDiffWithinAt 𝕜 n f s x theorem ContDiffWithinAt.fderivWithin_right_apply {f : F → G} {k : F → F} {s : Set F} {n : ℕ∞} {x₀ : F} (hf : ContDiffWithinAt 𝕜 n f s x₀) (hk : ContDiffWithinAt 𝕜 m k s x₀) - (hs : UniqueDiffOn 𝕜 s) (hmn : (m + 1 : ℕ∞) ≤ n) (hx₀s : x₀ ∈ s) : + (hs : UniqueDiffOn 𝕜 s) (hmn : m + 1 ≤ n) (hx₀s : x₀ ∈ s) : ContDiffWithinAt 𝕜 m (fun x => fderivWithin 𝕜 f s x (k x)) s x₀ := ContDiffWithinAt.fderivWithin_apply (ContDiffWithinAt.comp (x₀, x₀) hf contDiffWithinAt_snd <| prod_subset_preimage_snd s s) @@ -1044,10 +1048,10 @@ theorem ContDiffWithinAt.fderivWithin_right_apply -- TODO: can we make a version of `ContDiffWithinAt.fderivWithin` for iterated derivatives? theorem ContDiffWithinAt.iteratedFderivWithin_right {i : ℕ} (hf : ContDiffWithinAt 𝕜 n f s x₀) - (hs : UniqueDiffOn 𝕜 s) (hmn : (m + i : ℕ∞) ≤ n) (hx₀s : x₀ ∈ s) : + (hs : UniqueDiffOn 𝕜 s) (hmn : m + i ≤ n) (hx₀s : x₀ ∈ s) : ContDiffWithinAt 𝕜 m (iteratedFDerivWithin 𝕜 i f s) s x₀ := by induction' i with i hi generalizing m - · rw [ENat.coe_zero, add_zero] at hmn + · simp only [CharP.cast_eq_zero, add_zero] at hmn exact (hf.of_le hmn).continuousLinearMap_comp ((continuousMultilinearCurryFin0 𝕜 E F).symm : _ →L[𝕜] E [×0]→L[𝕜] F) · rw [Nat.cast_succ, add_comm _ 1, ← add_assoc] at hmn @@ -1056,7 +1060,7 @@ theorem ContDiffWithinAt.iteratedFderivWithin_right {i : ℕ} (hf : ContDiffWith _ →L[𝕜] E [×(i+1)]→L[𝕜] F) /-- `x ↦ fderiv 𝕜 (f x) (g x)` is smooth at `x₀`. -/ -protected theorem ContDiffAt.fderiv {f : E → F → G} {g : E → F} {n : ℕ∞} +protected theorem ContDiffAt.fderiv {f : E → F → G} {g : E → F} (hf : ContDiffAt 𝕜 n (Function.uncurry f) (x₀, g x₀)) (hg : ContDiffAt 𝕜 m g x₀) (hmn : m + 1 ≤ n) : ContDiffAt 𝕜 m (fun x => fderiv 𝕜 (f x) (g x)) x₀ := by simp_rw [← fderivWithin_univ] @@ -1065,44 +1069,44 @@ protected theorem ContDiffAt.fderiv {f : E → F → G} {g : E → F} {n : ℕ rw [preimage_univ] /-- `fderiv 𝕜 f` is smooth at `x₀`. -/ -theorem ContDiffAt.fderiv_right (hf : ContDiffAt 𝕜 n f x₀) (hmn : (m + 1 : ℕ∞) ≤ n) : +theorem ContDiffAt.fderiv_right (hf : ContDiffAt 𝕜 n f x₀) (hmn : m + 1 ≤ n) : ContDiffAt 𝕜 m (fderiv 𝕜 f) x₀ := ContDiffAt.fderiv (ContDiffAt.comp (x₀, x₀) hf contDiffAt_snd) contDiffAt_id hmn theorem ContDiffAt.iteratedFDeriv_right {i : ℕ} (hf : ContDiffAt 𝕜 n f x₀) - (hmn : (m + i : ℕ∞) ≤ n) : ContDiffAt 𝕜 m (iteratedFDeriv 𝕜 i f) x₀ := by + (hmn : m + i ≤ n) : ContDiffAt 𝕜 m (iteratedFDeriv 𝕜 i f) x₀ := by rw [← iteratedFDerivWithin_univ, ← contDiffWithinAt_univ] at * exact hf.iteratedFderivWithin_right uniqueDiffOn_univ hmn trivial /-- `x ↦ fderiv 𝕜 (f x) (g x)` is smooth. -/ -protected theorem ContDiff.fderiv {f : E → F → G} {g : E → F} {n m : ℕ∞} +protected theorem ContDiff.fderiv {f : E → F → G} {g : E → F} (hf : ContDiff 𝕜 m <| Function.uncurry f) (hg : ContDiff 𝕜 n g) (hnm : n + 1 ≤ m) : ContDiff 𝕜 n fun x => fderiv 𝕜 (f x) (g x) := contDiff_iff_contDiffAt.mpr fun _ => hf.contDiffAt.fderiv hg.contDiffAt hnm /-- `fderiv 𝕜 f` is smooth. -/ -theorem ContDiff.fderiv_right (hf : ContDiff 𝕜 n f) (hmn : (m + 1 : ℕ∞) ≤ n) : +theorem ContDiff.fderiv_right (hf : ContDiff 𝕜 n f) (hmn : m + 1 ≤ n) : ContDiff 𝕜 m (fderiv 𝕜 f) := contDiff_iff_contDiffAt.mpr fun _x => hf.contDiffAt.fderiv_right hmn theorem ContDiff.iteratedFDeriv_right {i : ℕ} (hf : ContDiff 𝕜 n f) - (hmn : (m + i : ℕ∞) ≤ n) : ContDiff 𝕜 m (iteratedFDeriv 𝕜 i f) := + (hmn : m + i ≤ n) : ContDiff 𝕜 m (iteratedFDeriv 𝕜 i f) := contDiff_iff_contDiffAt.mpr fun _x => hf.contDiffAt.iteratedFDeriv_right hmn /-- `x ↦ fderiv 𝕜 (f x) (g x)` is continuous. -/ -theorem Continuous.fderiv {f : E → F → G} {g : E → F} {n : ℕ∞} +theorem Continuous.fderiv {f : E → F → G} {g : E → F} (hf : ContDiff 𝕜 n <| Function.uncurry f) (hg : Continuous g) (hn : 1 ≤ n) : Continuous fun x => fderiv 𝕜 (f x) (g x) := (hf.fderiv (contDiff_zero.mpr hg) hn).continuous /-- `x ↦ fderiv 𝕜 (f x) (g x) (k x)` is smooth. -/ -theorem ContDiff.fderiv_apply {f : E → F → G} {g k : E → F} {n m : ℕ∞} +theorem ContDiff.fderiv_apply {f : E → F → G} {g k : E → F} (hf : ContDiff 𝕜 m <| Function.uncurry f) (hg : ContDiff 𝕜 n g) (hk : ContDiff 𝕜 n k) (hnm : n + 1 ≤ m) : ContDiff 𝕜 n fun x => fderiv 𝕜 (f x) (g x) (k x) := (hf.fderiv hg hnm).clm_apply hk /-- The bundled derivative of a `C^{n+1}` function is `C^n`. -/ -theorem contDiffOn_fderivWithin_apply {m n : ℕ∞} {s : Set E} {f : E → F} (hf : ContDiffOn 𝕜 n f s) +theorem contDiffOn_fderivWithin_apply {s : Set E} {f : E → F} (hf : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s) (hmn : m + 1 ≤ n) : ContDiffOn 𝕜 m (fun p : E × E => (fderivWithin 𝕜 f s p.1 : E →L[𝕜] F) p.2) (s ×ˢ univ) := ((hf.fderivWithin hs hmn).comp contDiffOn_fst (prod_subset_preimage_fst _ _)).clm_apply @@ -1176,7 +1180,7 @@ theorem contDiffAt_pi : ContDiffAt 𝕜 n Φ x ↔ ∀ i, ContDiffAt 𝕜 n (fun theorem contDiff_pi : ContDiff 𝕜 n Φ ↔ ∀ i, ContDiff 𝕜 n fun x => Φ x i := by simp only [← contDiffOn_univ, contDiffOn_pi] -theorem contDiff_update [DecidableEq ι] (k : ℕ∞) (x : ∀ i, F' i) (i : ι) : +theorem contDiff_update [DecidableEq ι] (k : WithTop ℕ∞) (x : ∀ i, F' i) (i : ι) : ContDiff 𝕜 k (update x i) := by rw [contDiff_pi] intro j @@ -1187,7 +1191,7 @@ theorem contDiff_update [DecidableEq ι] (k : ℕ∞) (x : ∀ i, F' i) (i : ι) · exact contDiff_const variable (F') in -theorem contDiff_single [DecidableEq ι] (k : ℕ∞) (i : ι) : +theorem contDiff_single [DecidableEq ι] (k : WithTop ℕ∞) (i : ι) : ContDiff 𝕜 k (Pi.single i : F' i → ∀ i, F' i) := contDiff_update k 0 i @@ -1207,7 +1211,7 @@ section Add theorem HasFTaylorSeriesUpToOn.add {n : WithTop ℕ∞} {q g} (hf : HasFTaylorSeriesUpToOn n f p s) (hg : HasFTaylorSeriesUpToOn n g q s) : HasFTaylorSeriesUpToOn n (f + g) (p + q) s := by - convert HasFTaylorSeriesUpToOn.continuousLinearMap_comp + exact HasFTaylorSeriesUpToOn.continuousLinearMap_comp (ContinuousLinearMap.fst 𝕜 F F + .snd 𝕜 F F) (hf.prod hg) -- The sum is smooth. @@ -1637,6 +1641,10 @@ invertible element. The proof is by induction, bootstrapping using an identity derivative of inversion as a bilinear map of inversion itself. -/ theorem contDiffAt_ring_inverse [CompleteSpace R] (x : Rˣ) : ContDiffAt 𝕜 n Ring.inverse (x : R) := by + suffices H : ∀ (n : ℕ∞), ContDiffAt 𝕜 n Ring.inverse (x : R) by + intro k hk + exact H ⊤ k (mod_cast le_top) + intro n induction' n using ENat.nat_induction with n IH Itop · intro m hm refine ⟨{ y : R | IsUnit y }, ?_, ?_⟩ @@ -1648,7 +1656,7 @@ theorem contDiffAt_ring_inverse [CompleteSpace R] (x : Rˣ) : · rintro _ ⟨x', rfl⟩ exact (inverse_continuousAt x').continuousWithinAt · simp [ftaylorSeriesWithin] - · rw [show (n.succ : ℕ∞) = n + 1 from rfl, contDiffAt_succ_iff_hasFDerivAt] + · rw [show ((n.succ : ℕ∞) : WithTop ℕ∞) = n + 1 from rfl, contDiffAt_succ_iff_hasFDerivAt] refine ⟨fun x : R => -mulLeftRight 𝕜 R (inverse x) (inverse x), ?_, ?_⟩ · refine ⟨{ y : R | IsUnit y }, x.nhds, ?_⟩ rintro _ ⟨y, rfl⟩ @@ -1671,36 +1679,36 @@ variable {𝕜} -- TODO: the next few lemmas don't need `𝕜` or `𝕜'` to be complete -- A good way to show this is to generalize `contDiffAt_ring_inverse` to the setting -- of a function `f` such that `∀ᶠ x in 𝓝 a, x * f x = 1`. -theorem ContDiffWithinAt.inv {f : E → 𝕜'} {n} (hf : ContDiffWithinAt 𝕜 n f s x) (hx : f x ≠ 0) : +theorem ContDiffWithinAt.inv {f : E → 𝕜'} (hf : ContDiffWithinAt 𝕜 n f s x) (hx : f x ≠ 0) : ContDiffWithinAt 𝕜 n (fun x => (f x)⁻¹) s x := (contDiffAt_inv 𝕜 hx).comp_contDiffWithinAt x hf -theorem ContDiffOn.inv {f : E → 𝕜'} {n} (hf : ContDiffOn 𝕜 n f s) (h : ∀ x ∈ s, f x ≠ 0) : +theorem ContDiffOn.inv {f : E → 𝕜'} (hf : ContDiffOn 𝕜 n f s) (h : ∀ x ∈ s, f x ≠ 0) : ContDiffOn 𝕜 n (fun x => (f x)⁻¹) s := fun x hx => (hf.contDiffWithinAt hx).inv (h x hx) -nonrec theorem ContDiffAt.inv {f : E → 𝕜'} {n} (hf : ContDiffAt 𝕜 n f x) (hx : f x ≠ 0) : +nonrec theorem ContDiffAt.inv {f : E → 𝕜'} (hf : ContDiffAt 𝕜 n f x) (hx : f x ≠ 0) : ContDiffAt 𝕜 n (fun x => (f x)⁻¹) x := hf.inv hx -theorem ContDiff.inv {f : E → 𝕜'} {n} (hf : ContDiff 𝕜 n f) (h : ∀ x, f x ≠ 0) : +theorem ContDiff.inv {f : E → 𝕜'} (hf : ContDiff 𝕜 n f) (h : ∀ x, f x ≠ 0) : ContDiff 𝕜 n fun x => (f x)⁻¹ := by rw [contDiff_iff_contDiffAt]; exact fun x => hf.contDiffAt.inv (h x) -- TODO: generalize to `f g : E → 𝕜'` -theorem ContDiffWithinAt.div [CompleteSpace 𝕜] {f g : E → 𝕜} {n} (hf : ContDiffWithinAt 𝕜 n f s x) +theorem ContDiffWithinAt.div [CompleteSpace 𝕜] {f g : E → 𝕜} (hf : ContDiffWithinAt 𝕜 n f s x) (hg : ContDiffWithinAt 𝕜 n g s x) (hx : g x ≠ 0) : ContDiffWithinAt 𝕜 n (fun x => f x / g x) s x := by simpa only [div_eq_mul_inv] using hf.mul (hg.inv hx) -theorem ContDiffOn.div [CompleteSpace 𝕜] {f g : E → 𝕜} {n} (hf : ContDiffOn 𝕜 n f s) +theorem ContDiffOn.div [CompleteSpace 𝕜] {f g : E → 𝕜} (hf : ContDiffOn 𝕜 n f s) (hg : ContDiffOn 𝕜 n g s) (h₀ : ∀ x ∈ s, g x ≠ 0) : ContDiffOn 𝕜 n (f / g) s := fun x hx => (hf x hx).div (hg x hx) (h₀ x hx) -nonrec theorem ContDiffAt.div [CompleteSpace 𝕜] {f g : E → 𝕜} {n} (hf : ContDiffAt 𝕜 n f x) +nonrec theorem ContDiffAt.div [CompleteSpace 𝕜] {f g : E → 𝕜} (hf : ContDiffAt 𝕜 n f x) (hg : ContDiffAt 𝕜 n g x) (hx : g x ≠ 0) : ContDiffAt 𝕜 n (fun x => f x / g x) x := hf.div hg hx -theorem ContDiff.div [CompleteSpace 𝕜] {f g : E → 𝕜} {n} (hf : ContDiff 𝕜 n f) (hg : ContDiff 𝕜 n g) +theorem ContDiff.div [CompleteSpace 𝕜] {f g : E → 𝕜} (hf : ContDiff 𝕜 n f) (hg : ContDiff 𝕜 n g) (h0 : ∀ x, g x ≠ 0) : ContDiff 𝕜 n fun x => f x / g x := by simp only [contDiff_iff_contDiffAt] at * exact fun x => (hf x).div (hg x) (h0 x) @@ -1738,23 +1746,17 @@ section FunctionInverse open ContinuousLinearMap -/-- If `f` is a local homeomorphism and the point `a` is in its target, -and if `f` is `n` times continuously differentiable at `f.symm a`, -and if the derivative at `f.symm a` is a continuous linear equivalence, -then `f.symm` is `n` times continuously differentiable at the point `a`. - -This is one of the easy parts of the inverse function theorem: it assumes that we already have -an inverse function. -/ -theorem PartialHomeomorph.contDiffAt_symm [CompleteSpace E] (f : PartialHomeomorph E F) +private theorem PartialHomeomorph.contDiffAt_symm_aux {n : ℕ∞} + [CompleteSpace E] (f : PartialHomeomorph E F) {f₀' : E ≃L[𝕜] F} {a : F} (ha : a ∈ f.target) (hf₀' : HasFDerivAt f (f₀' : E →L[𝕜] F) (f.symm a)) (hf : ContDiffAt 𝕜 n f (f.symm a)) : ContDiffAt 𝕜 n f.symm a := by - -- We prove this by induction on `n` + -- We prove this by induction on `n` induction' n using ENat.nat_induction with n IH Itop - · rw [contDiffAt_zero] + · apply contDiffAt_zero.2 exact ⟨f.target, IsOpen.mem_nhds f.open_target ha, f.continuousOn_invFun⟩ · obtain ⟨f', ⟨u, hu, hff'⟩, hf'⟩ := contDiffAt_succ_iff_hasFDerivAt.mp hf - rw [show (n.succ : ℕ∞) = n + 1 from rfl, contDiffAt_succ_iff_hasFDerivAt] + rw [show ((n.succ : ℕ∞) : WithTop ℕ∞) = n + 1 from rfl, contDiffAt_succ_iff_hasFDerivAt] -- For showing `n.succ` times continuous differentiability (the main inductive step), it -- suffices to produce the derivative and show that it is `n` times continuously differentiable have eq_f₀' : f' (f.symm a) = f₀' := (hff' (f.symm a) (mem_of_mem_nhds hu)).unique hf₀' @@ -1791,6 +1793,24 @@ theorem PartialHomeomorph.contDiffAt_symm [CompleteSpace E] (f : PartialHomeomor intro n exact Itop n (contDiffAt_top.mp hf n) +/-- If `f` is a local homeomorphism and the point `a` is in its target, +and if `f` is `n` times continuously differentiable at `f.symm a`, +and if the derivative at `f.symm a` is a continuous linear equivalence, +then `f.symm` is `n` times continuously differentiable at the point `a`. + +This is one of the easy parts of the inverse function theorem: it assumes that we already have +an inverse function. -/ +theorem PartialHomeomorph.contDiffAt_symm [CompleteSpace E] (f : PartialHomeomorph E F) + {f₀' : E ≃L[𝕜] F} {a : F} (ha : a ∈ f.target) + (hf₀' : HasFDerivAt f (f₀' : E →L[𝕜] F) (f.symm a)) (hf : ContDiffAt 𝕜 n f (f.symm a)) : + ContDiffAt 𝕜 n f.symm a := by + match n with + | ω => + intro k hk + exact f.contDiffAt_symm_aux ha hf₀' (hf.of_le (m := k) le_top) k le_rfl + | (n : ℕ∞) => + exact f.contDiffAt_symm_aux ha hf₀' hf + /-- If `f` is an `n` times continuously differentiable homeomorphism, and if the derivative of `f` at each point is a continuous linear equivalence, then `f.symm` is `n` times continuously differentiable. @@ -1905,14 +1925,14 @@ theorem contDiffOn_top_iff_derivWithin (hs : UniqueDiffOn 𝕜 s₂) : ContDiffOn 𝕜 ∞ f₂ s₂ ↔ DifferentiableOn 𝕜 f₂ s₂ ∧ ContDiffOn 𝕜 ∞ (derivWithin f₂ s₂) s₂ := by constructor · intro h - refine ⟨h.differentiableOn le_top, ?_⟩ + refine ⟨h.differentiableOn (mod_cast le_top), ?_⟩ refine contDiffOn_top.2 fun n => ((contDiffOn_succ_iff_derivWithin hs).1 ?_).2 - exact h.of_le le_top + exact h.of_le (mod_cast le_top) · intro h refine contDiffOn_top.2 fun n => ?_ - have A : (n : ℕ∞) ≤ ∞ := le_top + have A : (n : ℕ∞) ≤ ∞ := mod_cast le_top apply ((contDiffOn_succ_iff_derivWithin hs).2 ⟨h.1, h.2.of_le A⟩).of_le - exact WithTop.coe_le_coe.2 (Nat.le_succ n) + exact_mod_cast (Nat.le_succ n) /-- A function is `C^∞` on an open domain if and only if it is differentiable there, and its derivative (formulated with `deriv`) is `C^∞`. -/ @@ -1923,13 +1943,17 @@ theorem contDiffOn_top_iff_deriv_of_isOpen (hs : IsOpen s₂) : protected theorem ContDiffOn.derivWithin (hf : ContDiffOn 𝕜 n f₂ s₂) (hs : UniqueDiffOn 𝕜 s₂) (hmn : m + 1 ≤ n) : ContDiffOn 𝕜 m (derivWithin f₂ s₂) s₂ := by - cases m - · change ∞ + 1 ≤ n at hmn - have : n = ∞ := by simpa using hmn - rw [this] at hf - exact ((contDiffOn_top_iff_derivWithin hs).1 hf).2 - · change (Nat.succ _ : ℕ∞) ≤ n at hmn - exact ((contDiffOn_succ_iff_derivWithin hs).1 (hf.of_le hmn)).2 + rcases le_or_lt ∞ n with hn | hn + · have : ContDiffOn 𝕜 ∞ (derivWithin f₂ s₂) s₂ := + ((contDiffOn_top_iff_derivWithin hs).1 (hf.of_le hn)).2 + intro x hx k hk + exact this x hx k (mod_cast le_top) + · match m with + | ω => simpa using hmn.trans_lt hn + | ∞ => simpa using hmn.trans_lt hn + | (m : ℕ) => + change (m.succ : ℕ∞) ≤ n at hmn + exact ((contDiffOn_succ_iff_derivWithin hs).1 (hf.of_le hmn)).2 theorem ContDiffOn.deriv_of_isOpen (hf : ContDiffOn 𝕜 n f₂ s₂) (hs : IsOpen s₂) (hmn : m + 1 ≤ n) : ContDiffOn 𝕜 m (deriv f₂) s₂ := diff --git a/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean b/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean index 3d7d145291bee..fe52bff69b40d 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean @@ -53,7 +53,7 @@ theorem ContinuousLinearMap.norm_iteratedFDerivWithin_le_of_bilinear_aux {Du Eu · simp only [norm_iteratedFDerivWithin_zero, zero_add, Finset.range_one, Finset.sum_singleton, Nat.choose_self, Nat.cast_one, one_mul, Nat.sub_zero, ← mul_assoc] apply B.le_opNorm₂ - · have In : (n : ℕ∞) + 1 ≤ n.succ := by simp only [Nat.cast_succ, le_refl] + · have In : (n : WithTop ℕ∞) + 1 ≤ n.succ := by simp only [Nat.cast_succ, le_refl] -- Porting note: the next line is a hack allowing Lean to find the operator norm instance. let norm := @ContinuousLinearMap.hasOpNorm _ _ Eu ((Du →L[𝕜] Fu) →L[𝕜] Du →L[𝕜] Gu) _ _ _ _ _ _ (RingHom.id 𝕜) @@ -101,7 +101,7 @@ theorem ContinuousLinearMap.norm_iteratedFDerivWithin_le_of_bilinear_aux {Du Eu iteratedFDerivWithin 𝕜 n (fun y => B.precompR Du (f y) (fderivWithin 𝕜 g s y) + B.precompL Du (fderivWithin 𝕜 f s y) (g y)) s x := by apply iteratedFDerivWithin_congr (fun y hy => ?_) hx - have L : (1 : ℕ∞) ≤ n.succ := by + have L : (1 : WithTop ℕ∞) ≤ n.succ := by simpa only [ENat.coe_one, Nat.one_le_cast] using Nat.succ_pos n exact B.fderivWithin_of_bilinear (hf.differentiableOn L y hy) (hg.differentiableOn L y hy) (hs y hy) @@ -123,8 +123,8 @@ theorem ContinuousLinearMap.norm_iteratedFDerivWithin_le_of_bilinear_aux {Du Eu iterated derivatives of `f` and `g` when `B` is bilinear: `‖D^n (x ↦ B (f x) (g x))‖ ≤ ‖B‖ ∑_{k ≤ n} n.choose k ‖D^k f‖ ‖D^{n-k} g‖` -/ theorem ContinuousLinearMap.norm_iteratedFDerivWithin_le_of_bilinear (B : E →L[𝕜] F →L[𝕜] G) - {f : D → E} {g : D → F} {N : ℕ∞} {s : Set D} {x : D} (hf : ContDiffOn 𝕜 N f s) - (hg : ContDiffOn 𝕜 N g s) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {n : ℕ} (hn : (n : ℕ∞) ≤ N) : + {f : D → E} {g : D → F} {N : WithTop ℕ∞} {s : Set D} {x : D} (hf : ContDiffOn 𝕜 N f s) + (hg : ContDiffOn 𝕜 N g s) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {n : ℕ} (hn : n ≤ N) : ‖iteratedFDerivWithin 𝕜 n (fun y => B (f y) (g y)) s x‖ ≤ ‖B‖ * ∑ i ∈ Finset.range (n + 1), (n.choose i : ℝ) * ‖iteratedFDerivWithin 𝕜 i f s x‖ * ‖iteratedFDerivWithin 𝕜 (n - i) g s x‖ := by @@ -205,8 +205,8 @@ theorem ContinuousLinearMap.norm_iteratedFDerivWithin_le_of_bilinear (B : E →L iterated derivatives of `f` and `g` when `B` is bilinear: `‖D^n (x ↦ B (f x) (g x))‖ ≤ ‖B‖ ∑_{k ≤ n} n.choose k ‖D^k f‖ ‖D^{n-k} g‖` -/ theorem ContinuousLinearMap.norm_iteratedFDeriv_le_of_bilinear (B : E →L[𝕜] F →L[𝕜] G) {f : D → E} - {g : D → F} {N : ℕ∞} (hf : ContDiff 𝕜 N f) (hg : ContDiff 𝕜 N g) (x : D) {n : ℕ} - (hn : (n : ℕ∞) ≤ N) : + {g : D → F} {N : WithTop ℕ∞} (hf : ContDiff 𝕜 N f) (hg : ContDiff 𝕜 N g) (x : D) {n : ℕ} + (hn : n ≤ N) : ‖iteratedFDeriv 𝕜 n (fun y => B (f y) (g y)) x‖ ≤ ‖B‖ * ∑ i ∈ Finset.range (n + 1), (n.choose i : ℝ) * ‖iteratedFDeriv 𝕜 i f x‖ * ‖iteratedFDeriv 𝕜 (n - i) g x‖ := by simp_rw [← iteratedFDerivWithin_univ] @@ -217,9 +217,9 @@ theorem ContinuousLinearMap.norm_iteratedFDeriv_le_of_bilinear (B : E →L[𝕜] iterated derivatives of `f` and `g` when `B` is bilinear of norm at most `1`: `‖D^n (x ↦ B (f x) (g x))‖ ≤ ∑_{k ≤ n} n.choose k ‖D^k f‖ ‖D^{n-k} g‖` -/ theorem ContinuousLinearMap.norm_iteratedFDerivWithin_le_of_bilinear_of_le_one - (B : E →L[𝕜] F →L[𝕜] G) {f : D → E} {g : D → F} {N : ℕ∞} {s : Set D} {x : D} + (B : E →L[𝕜] F →L[𝕜] G) {f : D → E} {g : D → F} {N : WithTop ℕ∞} {s : Set D} {x : D} (hf : ContDiffOn 𝕜 N f s) (hg : ContDiffOn 𝕜 N g s) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {n : ℕ} - (hn : (n : ℕ∞) ≤ N) (hB : ‖B‖ ≤ 1) : ‖iteratedFDerivWithin 𝕜 n (fun y => B (f y) (g y)) s x‖ ≤ + (hn : n ≤ N) (hB : ‖B‖ ≤ 1) : ‖iteratedFDerivWithin 𝕜 n (fun y => B (f y) (g y)) s x‖ ≤ ∑ i ∈ Finset.range (n + 1), (n.choose i : ℝ) * ‖iteratedFDerivWithin 𝕜 i f s x‖ * ‖iteratedFDerivWithin 𝕜 (n - i) g s x‖ := by apply (B.norm_iteratedFDerivWithin_le_of_bilinear hf hg hs hx hn).trans @@ -229,8 +229,9 @@ theorem ContinuousLinearMap.norm_iteratedFDerivWithin_le_of_bilinear_of_le_one iterated derivatives of `f` and `g` when `B` is bilinear of norm at most `1`: `‖D^n (x ↦ B (f x) (g x))‖ ≤ ∑_{k ≤ n} n.choose k ‖D^k f‖ ‖D^{n-k} g‖` -/ theorem ContinuousLinearMap.norm_iteratedFDeriv_le_of_bilinear_of_le_one (B : E →L[𝕜] F →L[𝕜] G) - {f : D → E} {g : D → F} {N : ℕ∞} (hf : ContDiff 𝕜 N f) (hg : ContDiff 𝕜 N g) (x : D) {n : ℕ} - (hn : (n : ℕ∞) ≤ N) (hB : ‖B‖ ≤ 1) : ‖iteratedFDeriv 𝕜 n (fun y => B (f y) (g y)) x‖ ≤ + {f : D → E} {g : D → F} {N : WithTop ℕ∞} (hf : ContDiff 𝕜 N f) (hg : ContDiff 𝕜 N g) + (x : D) {n : ℕ} (hn : n ≤ N) (hB : ‖B‖ ≤ 1) : + ‖iteratedFDeriv 𝕜 n (fun y => B (f y) (g y)) x‖ ≤ ∑ i ∈ Finset.range (n + 1), (n.choose i : ℝ) * ‖iteratedFDeriv 𝕜 i f x‖ * ‖iteratedFDeriv 𝕜 (n - i) g x‖ := by simp_rw [← iteratedFDerivWithin_univ] @@ -242,17 +243,17 @@ section variable {𝕜' : Type*} [NormedField 𝕜'] [NormedAlgebra 𝕜 𝕜'] [NormedSpace 𝕜' F] [IsScalarTower 𝕜 𝕜' F] -theorem norm_iteratedFDerivWithin_smul_le {f : E → 𝕜'} {g : E → F} {N : ℕ∞} +theorem norm_iteratedFDerivWithin_smul_le {f : E → 𝕜'} {g : E → F} {N : WithTop ℕ∞} (hf : ContDiffOn 𝕜 N f s) (hg : ContDiffOn 𝕜 N g s) (hs : UniqueDiffOn 𝕜 s) {x : E} (hx : x ∈ s) - {n : ℕ} (hn : (n : ℕ∞) ≤ N) : ‖iteratedFDerivWithin 𝕜 n (fun y => f y • g y) s x‖ ≤ + {n : ℕ} (hn : n ≤ N) : ‖iteratedFDerivWithin 𝕜 n (fun y => f y • g y) s x‖ ≤ ∑ i ∈ Finset.range (n + 1), (n.choose i : ℝ) * ‖iteratedFDerivWithin 𝕜 i f s x‖ * ‖iteratedFDerivWithin 𝕜 (n - i) g s x‖ := (ContinuousLinearMap.lsmul 𝕜 𝕜' : 𝕜' →L[𝕜] F →L[𝕜] F).norm_iteratedFDerivWithin_le_of_bilinear_of_le_one hf hg hs hx hn ContinuousLinearMap.opNorm_lsmul_le -theorem norm_iteratedFDeriv_smul_le {f : E → 𝕜'} {g : E → F} {N : ℕ∞} (hf : ContDiff 𝕜 N f) - (hg : ContDiff 𝕜 N g) (x : E) {n : ℕ} (hn : (n : ℕ∞) ≤ N) : +theorem norm_iteratedFDeriv_smul_le {f : E → 𝕜'} {g : E → F} {N : WithTop ℕ∞} (hf : ContDiff 𝕜 N f) + (hg : ContDiff 𝕜 N g) (x : E) {n : ℕ} (hn : n ≤ N) : ‖iteratedFDeriv 𝕜 n (fun y => f y • g y) x‖ ≤ ∑ i ∈ Finset.range (n + 1), (n.choose i : ℝ) * ‖iteratedFDeriv 𝕜 i f x‖ * ‖iteratedFDeriv 𝕜 (n - i) g x‖ := (ContinuousLinearMap.lsmul 𝕜 𝕜' : 𝕜' →L[𝕜] F →L[𝕜] F).norm_iteratedFDeriv_le_of_bilinear_of_le_one @@ -265,17 +266,18 @@ section variable {ι : Type*} {A : Type*} [NormedRing A] [NormedAlgebra 𝕜 A] {A' : Type*} [NormedCommRing A'] [NormedAlgebra 𝕜 A'] -theorem norm_iteratedFDerivWithin_mul_le {f : E → A} {g : E → A} {N : ℕ∞} (hf : ContDiffOn 𝕜 N f s) - (hg : ContDiffOn 𝕜 N g s) (hs : UniqueDiffOn 𝕜 s) {x : E} (hx : x ∈ s) {n : ℕ} - (hn : (n : ℕ∞) ≤ N) : ‖iteratedFDerivWithin 𝕜 n (fun y => f y * g y) s x‖ ≤ +theorem norm_iteratedFDerivWithin_mul_le {f : E → A} {g : E → A} {N : WithTop ℕ∞} + (hf : ContDiffOn 𝕜 N f s) (hg : ContDiffOn 𝕜 N g s) (hs : UniqueDiffOn 𝕜 s) + {x : E} (hx : x ∈ s) {n : ℕ} (hn : n ≤ N) : + ‖iteratedFDerivWithin 𝕜 n (fun y => f y * g y) s x‖ ≤ ∑ i ∈ Finset.range (n + 1), (n.choose i : ℝ) * ‖iteratedFDerivWithin 𝕜 i f s x‖ * ‖iteratedFDerivWithin 𝕜 (n - i) g s x‖ := (ContinuousLinearMap.mul 𝕜 A : A →L[𝕜] A →L[𝕜] A).norm_iteratedFDerivWithin_le_of_bilinear_of_le_one hf hg hs hx hn (ContinuousLinearMap.opNorm_mul_le _ _) -theorem norm_iteratedFDeriv_mul_le {f : E → A} {g : E → A} {N : ℕ∞} (hf : ContDiff 𝕜 N f) - (hg : ContDiff 𝕜 N g) (x : E) {n : ℕ} (hn : (n : ℕ∞) ≤ N) : +theorem norm_iteratedFDeriv_mul_le {f : E → A} {g : E → A} {N : WithTop ℕ∞} (hf : ContDiff 𝕜 N f) + (hg : ContDiff 𝕜 N g) (x : E) {n : ℕ} (hn : n ≤ N) : ‖iteratedFDeriv 𝕜 n (fun y => f y * g y) x‖ ≤ ∑ i ∈ Finset.range (n + 1), (n.choose i : ℝ) * ‖iteratedFDeriv 𝕜 i f x‖ * ‖iteratedFDeriv 𝕜 (n - i) g x‖ := by simp_rw [← iteratedFDerivWithin_univ] @@ -285,8 +287,8 @@ theorem norm_iteratedFDeriv_mul_le {f : E → A} {g : E → A} {N : ℕ∞} (hf -- TODO: Add `norm_iteratedFDeriv[Within]_list_prod_le` for non-commutative `NormedRing A`. theorem norm_iteratedFDerivWithin_prod_le [DecidableEq ι] [NormOneClass A'] {u : Finset ι} - {f : ι → E → A'} {N : ℕ∞} (hf : ∀ i ∈ u, ContDiffOn 𝕜 N (f i) s) (hs : UniqueDiffOn 𝕜 s) {x : E} - (hx : x ∈ s) {n : ℕ} (hn : (n : ℕ∞) ≤ N) : + {f : ι → E → A'} {N : WithTop ℕ∞} (hf : ∀ i ∈ u, ContDiffOn 𝕜 N (f i) s) + (hs : UniqueDiffOn 𝕜 s) {x : E} (hx : x ∈ s) {n : ℕ} (hn : n ≤ N) : ‖iteratedFDerivWithin 𝕜 n (∏ j ∈ u, f j ·) s x‖ ≤ ∑ p ∈ u.sym n, (p : Multiset ι).multinomial * ∏ j ∈ u, ‖iteratedFDerivWithin 𝕜 (Multiset.count j p) (f j) s x‖ := by @@ -309,7 +311,7 @@ theorem norm_iteratedFDerivWithin_prod_le [DecidableEq ι] [NormOneClass A'] {u simp only [comp_apply, Finset.symInsertEquiv_symm_apply_coe] refine Finset.sum_le_sum ?_ intro m _ - specialize IH hf.2 (n := n - m) (le_trans (WithTop.coe_le_coe.mpr (n.sub_le m)) hn) + specialize IH hf.2 (n := n - m) (le_trans (by exact_mod_cast n.sub_le m) hn) refine le_trans (mul_le_mul_of_nonneg_left IH (by simp [mul_nonneg])) ?_ rw [Finset.mul_sum, ← Finset.sum_coe_sort] refine Finset.sum_le_sum ?_ @@ -329,8 +331,8 @@ theorem norm_iteratedFDerivWithin_prod_le [DecidableEq ι] [NormOneClass A'] {u rw [Sym.count_coe_fill_of_ne hji] theorem norm_iteratedFDeriv_prod_le [DecidableEq ι] [NormOneClass A'] {u : Finset ι} - {f : ι → E → A'} {N : ℕ∞} (hf : ∀ i ∈ u, ContDiff 𝕜 N (f i)) {x : E} {n : ℕ} - (hn : (n : ℕ∞) ≤ N) : + {f : ι → E → A'} {N : WithTop ℕ∞} (hf : ∀ i ∈ u, ContDiff 𝕜 N (f i)) {x : E} {n : ℕ} + (hn : n ≤ N) : ‖iteratedFDeriv 𝕜 n (∏ j ∈ u, f j ·) x‖ ≤ ∑ p ∈ u.sym n, (p : Multiset ι).multinomial * ∏ j ∈ u, ‖iteratedFDeriv 𝕜 ((p : Multiset ι).count j) (f j) x‖ := by @@ -362,7 +364,7 @@ theorem norm_iteratedFDerivWithin_comp_le_aux {Fu Gu : Type u} [NormedAddCommGro induction' n using Nat.case_strong_induction_on with n IH generalizing Gu · simpa [norm_iteratedFDerivWithin_zero, Nat.factorial_zero, algebraMap.coe_one, one_mul, pow_zero, mul_one, comp_apply] using hC 0 le_rfl - have M : (n : ℕ∞) < n.succ := Nat.cast_lt.2 n.lt_succ_self + have M : (n : WithTop ℕ∞) < n.succ := Nat.cast_lt.2 n.lt_succ_self have Cnonneg : 0 ≤ C := (norm_nonneg _).trans (hC 0 bot_le) have Dnonneg : 0 ≤ D := by have : 1 ≤ n + 1 := by simp only [le_add_iff_nonneg_left, zero_le'] @@ -404,7 +406,8 @@ theorem norm_iteratedFDerivWithin_comp_le_aux {Fu Gu : Type u} [NormedAddCommGro LinearIsometryEquiv.norm_map] _ = ‖iteratedFDerivWithin 𝕜 n (fun y : E => ContinuousLinearMap.compL 𝕜 E Fu Gu (fderivWithin 𝕜 g t (f y)) (fderivWithin 𝕜 f s y)) s x‖ := by - have L : (1 : ℕ∞) ≤ n.succ := by simpa only [ENat.coe_one, Nat.one_le_cast] using n.succ_pos + have L : (1 : WithTop ℕ∞) ≤ n.succ := by + simpa only [ENat.coe_one, Nat.one_le_cast] using n.succ_pos congr 1 refine iteratedFDerivWithin_congr (fun y hy => ?_) hx _ apply fderivWithin_comp _ _ _ hst (hs y hy) @@ -459,7 +462,7 @@ theorem norm_iteratedFDerivWithin_comp_le_aux {Fu Gu : Type u} [NormedAddCommGro within a set of `f` at `x` is bounded by `D^i` for all `1 ≤ i ≤ n`, then the `n`-th derivative of `g ∘ f` is bounded by `n! * C * D^n`. -/ theorem norm_iteratedFDerivWithin_comp_le {g : F → G} {f : E → F} {n : ℕ} {s : Set E} {t : Set F} - {x : E} {N : ℕ∞} (hg : ContDiffOn 𝕜 N g t) (hf : ContDiffOn 𝕜 N f s) (hn : (n : ℕ∞) ≤ N) + {x : E} {N : WithTop ℕ∞} (hg : ContDiffOn 𝕜 N g t) (hf : ContDiffOn 𝕜 N f s) (hn : n ≤ N) (ht : UniqueDiffOn 𝕜 t) (hs : UniqueDiffOn 𝕜 s) (hst : MapsTo f s t) (hx : x ∈ s) {C : ℝ} {D : ℝ} (hC : ∀ i, i ≤ n → ‖iteratedFDerivWithin 𝕜 i g t (f x)‖ ≤ C) (hD : ∀ i, 1 ≤ i → i ≤ n → ‖iteratedFDerivWithin 𝕜 i f s x‖ ≤ D ^ i) : @@ -509,8 +512,8 @@ theorem norm_iteratedFDerivWithin_comp_le {g : F → G} {f : E → F} {n : ℕ} /-- If the derivatives of `g` at `f x` are bounded by `C`, and the `i`-th derivative of `f` at `x` is bounded by `D^i` for all `1 ≤ i ≤ n`, then the `n`-th derivative of `g ∘ f` is bounded by `n! * C * D^n`. -/ -theorem norm_iteratedFDeriv_comp_le {g : F → G} {f : E → F} {n : ℕ} {N : ℕ∞} (hg : ContDiff 𝕜 N g) - (hf : ContDiff 𝕜 N f) (hn : (n : ℕ∞) ≤ N) (x : E) {C : ℝ} {D : ℝ} +theorem norm_iteratedFDeriv_comp_le {g : F → G} {f : E → F} {n : ℕ} {N : WithTop ℕ∞} + (hg : ContDiff 𝕜 N g) (hf : ContDiff 𝕜 N f) (hn : n ≤ N) (x : E) {C : ℝ} {D : ℝ} (hC : ∀ i, i ≤ n → ‖iteratedFDeriv 𝕜 i g (f x)‖ ≤ C) (hD : ∀ i, 1 ≤ i → i ≤ n → ‖iteratedFDeriv 𝕜 i f x‖ ≤ D ^ i) : ‖iteratedFDeriv 𝕜 n (g ∘ f) x‖ ≤ n ! * C * D ^ n := by @@ -521,8 +524,9 @@ theorem norm_iteratedFDeriv_comp_le {g : F → G} {f : E → F} {n : ℕ} {N : section Apply theorem norm_iteratedFDerivWithin_clm_apply {f : E → F →L[𝕜] G} {g : E → F} {s : Set E} {x : E} - {N : ℕ∞} {n : ℕ} (hf : ContDiffOn 𝕜 N f s) (hg : ContDiffOn 𝕜 N g s) (hs : UniqueDiffOn 𝕜 s) - (hx : x ∈ s) (hn : ↑n ≤ N) : ‖iteratedFDerivWithin 𝕜 n (fun y => (f y) (g y)) s x‖ ≤ + {N : WithTop ℕ∞} {n : ℕ} (hf : ContDiffOn 𝕜 N f s) (hg : ContDiffOn 𝕜 N g s) + (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) (hn : n ≤ N) : + ‖iteratedFDerivWithin 𝕜 n (fun y => (f y) (g y)) s x‖ ≤ ∑ i ∈ Finset.range (n + 1), ↑(n.choose i) * ‖iteratedFDerivWithin 𝕜 i f s x‖ * ‖iteratedFDerivWithin 𝕜 (n - i) g s x‖ := by let B : (F →L[𝕜] G) →L[𝕜] F →L[𝕜] G := ContinuousLinearMap.flip (ContinuousLinearMap.apply 𝕜 G) @@ -533,8 +537,8 @@ theorem norm_iteratedFDerivWithin_clm_apply {f : E → F →L[𝕜] G} {g : E rfl exact B.norm_iteratedFDerivWithin_le_of_bilinear_of_le_one hf hg hs hx hn hB -theorem norm_iteratedFDeriv_clm_apply {f : E → F →L[𝕜] G} {g : E → F} {N : ℕ∞} {n : ℕ} - (hf : ContDiff 𝕜 N f) (hg : ContDiff 𝕜 N g) (x : E) (hn : ↑n ≤ N) : +theorem norm_iteratedFDeriv_clm_apply {f : E → F →L[𝕜] G} {g : E → F} {N : WithTop ℕ∞} {n : ℕ} + (hf : ContDiff 𝕜 N f) (hg : ContDiff 𝕜 N g) (x : E) (hn : n ≤ N) : ‖iteratedFDeriv 𝕜 n (fun y : E => (f y) (g y)) x‖ ≤ ∑ i ∈ Finset.range (n + 1), ↑(n.choose i) * ‖iteratedFDeriv 𝕜 i f x‖ * ‖iteratedFDeriv 𝕜 (n - i) g x‖ := by simp only [← iteratedFDerivWithin_univ] @@ -542,7 +546,8 @@ theorem norm_iteratedFDeriv_clm_apply {f : E → F →L[𝕜] G} {g : E → F} { (Set.mem_univ x) hn theorem norm_iteratedFDerivWithin_clm_apply_const {f : E → F →L[𝕜] G} {c : F} {s : Set E} {x : E} - {N : ℕ∞} {n : ℕ} (hf : ContDiffOn 𝕜 N f s) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) (hn : ↑n ≤ N) : + {N : WithTop ℕ∞} {n : ℕ} (hf : ContDiffOn 𝕜 N f s) (hs : UniqueDiffOn 𝕜 s) + (hx : x ∈ s) (hn : n ≤ N) : ‖iteratedFDerivWithin 𝕜 n (fun y : E => (f y) c) s x‖ ≤ ‖c‖ * ‖iteratedFDerivWithin 𝕜 n f s x‖ := by let g : (F →L[𝕜] G) →L[𝕜] G := ContinuousLinearMap.apply 𝕜 G c @@ -553,8 +558,8 @@ theorem norm_iteratedFDerivWithin_clm_apply_const {f : E → F →L[𝕜] G} {c rw [ContinuousLinearMap.apply_apply, mul_comm] exact f.le_opNorm c -theorem norm_iteratedFDeriv_clm_apply_const {f : E → F →L[𝕜] G} {c : F} {x : E} {N : ℕ∞} {n : ℕ} - (hf : ContDiff 𝕜 N f) (hn : ↑n ≤ N) : +theorem norm_iteratedFDeriv_clm_apply_const {f : E → F →L[𝕜] G} {c : F} {x : E} + {N : WithTop ℕ∞} {n : ℕ} (hf : ContDiff 𝕜 N f) (hn : n ≤ N) : ‖iteratedFDeriv 𝕜 n (fun y : E => (f y) c) x‖ ≤ ‖c‖ * ‖iteratedFDeriv 𝕜 n f x‖ := by simp only [← iteratedFDerivWithin_univ] exact norm_iteratedFDerivWithin_clm_apply_const hf.contDiffOn uniqueDiffOn_univ diff --git a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean index 968b9b7d7d1c3..180ea22be1a5e 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean @@ -93,7 +93,13 @@ noncomputable section open scoped Classical open NNReal Topology Filter -local notation "∞" => (⊤ : ℕ∞) +/-- Smoothness exponent for analytic functions. Not implemented yet, `ω` smoothness is equivalent +to `∞` smoothness in current mathlib. -/ +scoped [ContDiff] notation3 "ω" => (⊤ : WithTop ℕ∞) +/-- Smoothness exponent for infinitely differentiable functions. -/ +scoped [ContDiff] notation3 "∞" => ((⊤ : ℕ∞) : WithTop ℕ∞) + +open ContDiff /- Porting note: These lines are not required in Mathlib4. @@ -108,7 +114,7 @@ universe u uE uF uG uX variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] {E : Type uE} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {F : Type uF} [NormedAddCommGroup F] [NormedSpace 𝕜 F] {G : Type uG} [NormedAddCommGroup G] [NormedSpace 𝕜 G] {X : Type uX} [NormedAddCommGroup X] [NormedSpace 𝕜 X] - {s s₁ t u : Set E} {f f₁ : E → F} {g : F → G} {x x₀ : E} {c : F} {m n : ℕ∞} + {s s₁ t u : Set E} {f f₁ : E → F} {g : F → G} {x x₀ : E} {c : F} {m n : WithTop ℕ∞} {p : E → FormalMultilinearSeries 𝕜 E F} /-! ### Smooth functions within a set around a point -/ @@ -122,8 +128,11 @@ depend on the finite order we consider). For instance, a real function which is `C^m` on `(-1/m, 1/m)` for each natural `m`, but not better, is `C^∞` at `0` within `univ`. + +We take the exponent `n` in `WithTop ℕ∞` to allow for an extension to analytic functions in the +future, but currently the notion is the same for `n = ∞` and `n = ω`. -/ -def ContDiffWithinAt (n : ℕ∞) (f : E → F) (s : Set E) (x : E) : Prop := +def ContDiffWithinAt (n : WithTop ℕ∞) (f : E → F) (s : Set E) (x : E) : Prop := ∀ m : ℕ, m ≤ n → ∃ u ∈ 𝓝[insert x s] x, ∃ p : E → FormalMultilinearSeries 𝕜 E F, HasFTaylorSeriesUpToOn m f p u @@ -137,9 +146,9 @@ theorem contDiffWithinAt_nat {n : ℕ} : theorem ContDiffWithinAt.of_le (h : ContDiffWithinAt 𝕜 n f s x) (hmn : m ≤ n) : ContDiffWithinAt 𝕜 m f s x := fun k hk => h k (le_trans hk hmn) -theorem contDiffWithinAt_iff_forall_nat_le : +theorem contDiffWithinAt_iff_forall_nat_le {n : ℕ∞} : ContDiffWithinAt 𝕜 n f s x ↔ ∀ m : ℕ, ↑m ≤ n → ContDiffWithinAt 𝕜 m f s x := - ⟨fun H _m hm => H.of_le hm, fun H m hm => H m hm _ le_rfl⟩ + ⟨fun H _m hm => H.of_le (mod_cast hm), fun H m hm => H m (mod_cast hm) _ le_rfl⟩ theorem contDiffWithinAt_top : ContDiffWithinAt 𝕜 ∞ f s x ↔ ∀ n : ℕ, ContDiffWithinAt 𝕜 n f s x := contDiffWithinAt_iff_forall_nat_le.trans <| by simp only [forall_prop_of_true, le_top] @@ -308,7 +317,7 @@ theorem contDiffWithinAt_succ_iff_hasFDerivWithinAt {n : ℕ} : hasFTaylorSeriesUpToOn_succ_iff_right] at Hp exact Hp.2.2.of_le (mod_cast hm) · rintro ⟨u, hu, f', f'_eq_deriv, Hf'⟩ - rw [show (n : ℕ∞) + 1 = (n + 1 : ℕ) from rfl, contDiffWithinAt_nat] + rw [show (n : WithTop ℕ∞) + 1 = (n + 1 : ℕ) from rfl, contDiffWithinAt_nat] rcases Hf' n le_rfl with ⟨v, hv, p', Hp'⟩ refine ⟨v ∩ u, ?_, fun x => (p' x).unshift (f x), ?_⟩ · apply Filter.inter_mem _ hu @@ -370,7 +379,7 @@ admits continuous derivatives up to order `n` on a neighborhood of `x` in `s`. For `n = ∞`, we only require that this holds up to any finite order (where the neighborhood may depend on the finite order we consider). -/ -def ContDiffOn (n : ℕ∞) (f : E → F) (s : Set E) : Prop := +def ContDiffOn (n : WithTop ℕ∞) (f : E → F) (s : Set E) : Prop := ∀ x ∈ s, ContDiffWithinAt 𝕜 n f s x variable {𝕜} @@ -427,13 +436,15 @@ theorem ContDiffOn.of_succ {n : ℕ} (h : ContDiffOn 𝕜 (n + 1) f s) : ContDif theorem ContDiffOn.one_of_succ {n : ℕ} (h : ContDiffOn 𝕜 (n + 1) f s) : ContDiffOn 𝕜 1 f s := h.of_le <| WithTop.coe_le_coe.mpr le_add_self -theorem contDiffOn_iff_forall_nat_le : ContDiffOn 𝕜 n f s ↔ ∀ m : ℕ, ↑m ≤ n → ContDiffOn 𝕜 m f s := - ⟨fun H _ hm => H.of_le hm, fun H x hx m hm => H m hm x hx m le_rfl⟩ +theorem contDiffOn_iff_forall_nat_le {n : ℕ∞} : + ContDiffOn 𝕜 n f s ↔ ∀ m : ℕ, ↑m ≤ n → ContDiffOn 𝕜 m f s := + ⟨fun H _ hm => H.of_le (mod_cast hm), fun H x hx m hm => H m (mod_cast hm) x hx m le_rfl⟩ theorem contDiffOn_top : ContDiffOn 𝕜 ∞ f s ↔ ∀ n : ℕ, ContDiffOn 𝕜 n f s := contDiffOn_iff_forall_nat_le.trans <| by simp only [le_top, forall_prop_of_true] -theorem contDiffOn_all_iff_nat : (∀ n, ContDiffOn 𝕜 n f s) ↔ ∀ n : ℕ, ContDiffOn 𝕜 n f s := by +theorem contDiffOn_all_iff_nat : + (∀ (n : ℕ∞), ContDiffOn 𝕜 n f s) ↔ ∀ n : ℕ, ContDiffOn 𝕜 n f s := by refine ⟨fun H n => H n, ?_⟩ rintro H (_ | n) exacts [contDiffOn_top.2 H, H n] @@ -522,8 +533,8 @@ protected theorem ContDiffOn.ftaylorSeriesWithin (h : ContDiffOn 𝕜 n f s) (hs simp only [ftaylorSeriesWithin, ContinuousMultilinearMap.curry0_apply, iteratedFDerivWithin_zero_apply] · intro m hm x hx - have : m < n := mod_cast hm - rcases (h x hx) m.succ (Order.add_one_le_of_lt this) with ⟨u, hu, p, Hp⟩ + have : (m + 1 : ℕ) ≤ n := ENat.add_one_natCast_le_withTop_of_lt hm + rcases (h x hx).of_le this _ le_rfl with ⟨u, hu, p, Hp⟩ rw [insert_eq_of_mem hx] at hu rcases mem_nhdsWithin.1 hu with ⟨o, o_open, xo, ho⟩ rw [inter_comm] at ho @@ -557,7 +568,7 @@ protected theorem ContDiffOn.ftaylorSeriesWithin (h : ContDiffOn 𝕜 n f s) (hs exact (Hp.mono ho).eq_iteratedFDerivWithin_of_uniqueDiffOn le_rfl (hs.inter o_open) ⟨hy, yo⟩ exact ((Hp.mono ho).cont m le_rfl).congr fun y hy => (A y hy).symm -theorem contDiffOn_of_continuousOn_differentiableOn +theorem contDiffOn_of_continuousOn_differentiableOn {n : ℕ∞} (Hcont : ∀ m : ℕ, m ≤ n → ContinuousOn (fun x => iteratedFDerivWithin 𝕜 m f s x) s) (Hdiff : ∀ m : ℕ, m < n → DifferentiableOn 𝕜 (fun x => iteratedFDerivWithin 𝕜 m f s x) s) : @@ -570,11 +581,11 @@ theorem contDiffOn_of_continuousOn_differentiableOn simp only [ftaylorSeriesWithin, ContinuousMultilinearMap.curry0_apply, iteratedFDerivWithin_zero_apply] · intro k hk y hy - convert (Hdiff k (lt_of_lt_of_le (mod_cast hk) hm) y hy).hasFDerivWithinAt + convert (Hdiff k (lt_of_lt_of_le (mod_cast hk) (mod_cast hm)) y hy).hasFDerivWithinAt · intro k hk - exact Hcont k (le_trans (mod_cast hk) hm) + exact Hcont k (le_trans (mod_cast hk) (mod_cast hm)) -theorem contDiffOn_of_differentiableOn +theorem contDiffOn_of_differentiableOn {n : ℕ∞} (h : ∀ m : ℕ, m ≤ n → DifferentiableOn 𝕜 (iteratedFDerivWithin 𝕜 m f s) s) : ContDiffOn 𝕜 n f s := contDiffOn_of_continuousOn_differentiableOn (fun m hm => (h m hm).continuousOn) fun m hm => @@ -592,7 +603,7 @@ theorem ContDiffOn.differentiableOn_iteratedFDerivWithin {m : ℕ} (h : ContDiff theorem ContDiffWithinAt.differentiableWithinAt_iteratedFDerivWithin {m : ℕ} (h : ContDiffWithinAt 𝕜 n f s x) (hmn : m < n) (hs : UniqueDiffOn 𝕜 (insert x s)) : DifferentiableWithinAt 𝕜 (iteratedFDerivWithin 𝕜 m f s) s x := by - rcases h.contDiffOn' (Order.add_one_le_of_lt hmn) with ⟨u, uo, xu, hu⟩ + rcases h.contDiffOn' (ENat.add_one_natCast_le_withTop_of_lt hmn) with ⟨u, uo, xu, hu⟩ set t := insert x s ∩ u have A : t =ᶠ[𝓝[≠] x] s := by simp only [set_eventuallyEq_iff_inf_principal, ← nhdsWithin_inter'] @@ -607,12 +618,12 @@ theorem ContDiffWithinAt.differentiableWithinAt_iteratedFDerivWithin {m : ℕ} rw [differentiableWithinAt_congr_set' _ A] at C exact C.congr_of_eventuallyEq (B.filter_mono inf_le_left) B.self_of_nhds -theorem contDiffOn_iff_continuousOn_differentiableOn (hs : UniqueDiffOn 𝕜 s) : +theorem contDiffOn_iff_continuousOn_differentiableOn {n : ℕ∞} (hs : UniqueDiffOn 𝕜 s) : ContDiffOn 𝕜 n f s ↔ (∀ m : ℕ, m ≤ n → ContinuousOn (fun x => iteratedFDerivWithin 𝕜 m f s x) s) ∧ ∀ m : ℕ, m < n → DifferentiableOn 𝕜 (fun x => iteratedFDerivWithin 𝕜 m f s x) s := - ⟨fun h => ⟨fun _m hm => h.continuousOn_iteratedFDerivWithin hm hs, fun _m hm => - h.differentiableOn_iteratedFDerivWithin hm hs⟩, + ⟨fun h => ⟨fun _m hm => h.continuousOn_iteratedFDerivWithin (mod_cast hm) hs, fun _m hm => + h.differentiableOn_iteratedFDerivWithin (mod_cast hm) hs⟩, fun h => contDiffOn_of_continuousOn_differentiableOn h.1 h.2⟩ theorem contDiffOn_succ_of_fderivWithin {n : ℕ} (hf : DifferentiableOn 𝕜 f s) @@ -628,7 +639,7 @@ theorem contDiffOn_succ_iff_fderivWithin {n : ℕ} (hs : UniqueDiffOn 𝕜 s) : ContDiffOn 𝕜 (n + 1) f s ↔ DifferentiableOn 𝕜 f s ∧ ContDiffOn 𝕜 n (fun y => fderivWithin 𝕜 f s y) s := by refine ⟨fun H => ?_, fun h => contDiffOn_succ_of_fderivWithin h.1 h.2⟩ - refine ⟨H.differentiableOn (WithTop.coe_le_coe.2 (Nat.le_add_left 1 n)), fun x hx => ?_⟩ + refine ⟨H.differentiableOn le_add_self, fun x hx => ?_⟩ rcases contDiffWithinAt_succ_iff_hasFDerivWithinAt.1 (H x hx) with ⟨u, hu, f', hff', hf'⟩ rcases mem_nhdsWithin.1 hu with ⟨o, o_open, xo, ho⟩ rw [inter_comm, insert_eq_of_mem hx] at ho @@ -666,14 +677,14 @@ theorem contDiffOn_top_iff_fderivWithin (hs : UniqueDiffOn 𝕜 s) : DifferentiableOn 𝕜 f s ∧ ContDiffOn 𝕜 ∞ (fun y => fderivWithin 𝕜 f s y) s := by constructor · intro h - refine ⟨h.differentiableOn le_top, ?_⟩ + refine ⟨h.differentiableOn (mod_cast le_top), ?_⟩ refine contDiffOn_top.2 fun n => ((contDiffOn_succ_iff_fderivWithin hs).1 ?_).2 - exact h.of_le le_top + exact h.of_le (mod_cast le_top) · intro h refine contDiffOn_top.2 fun n => ?_ - have A : (n : ℕ∞) ≤ ∞ := le_top + have A : (n : ℕ∞) ≤ ∞ := mod_cast le_top apply ((contDiffOn_succ_iff_fderivWithin hs).2 ⟨h.1, h.2.of_le A⟩).of_le - exact WithTop.coe_le_coe.2 (Nat.le_succ n) + exact_mod_cast (Nat.le_succ n) /-- A function is `C^∞` on an open domain if and only if it is differentiable there, and its derivative (expressed with `fderiv`) is `C^∞`. -/ @@ -684,13 +695,17 @@ theorem contDiffOn_top_iff_fderiv_of_isOpen (hs : IsOpen s) : protected theorem ContDiffOn.fderivWithin (hf : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s) (hmn : m + 1 ≤ n) : ContDiffOn 𝕜 m (fun y => fderivWithin 𝕜 f s y) s := by - cases' m with m - · change ∞ + 1 ≤ n at hmn - have : n = ∞ := by simpa using hmn - rw [this] at hf - exact ((contDiffOn_top_iff_fderivWithin hs).1 hf).2 - · change (m.succ : ℕ∞) ≤ n at hmn - exact ((contDiffOn_succ_iff_fderivWithin hs).1 (hf.of_le hmn)).2 + rcases le_or_lt ∞ n with hn | hn + · have : ContDiffOn 𝕜 ∞ (fun y ↦ fderivWithin 𝕜 f s y) s := + ((contDiffOn_top_iff_fderivWithin hs).1 (hf.of_le hn)).2 + intro x hx k hk + exact this x hx k (mod_cast le_top) + · match m with + | ω => simpa using hmn.trans_lt hn + | ∞ => simpa using hmn.trans_lt hn + | (m : ℕ) => + change (m.succ : ℕ∞) ≤ n at hmn + exact ((contDiffOn_succ_iff_fderivWithin hs).1 (hf.of_le hmn)).2 theorem ContDiffOn.fderiv_of_isOpen (hf : ContDiffOn 𝕜 n f s) (hs : IsOpen s) (hmn : m + 1 ≤ n) : ContDiffOn 𝕜 m (fun y => fderiv 𝕜 f y) s := @@ -704,6 +719,13 @@ theorem ContDiffOn.continuousOn_fderiv_of_isOpen (h : ContDiffOn 𝕜 n f s) (hs (hn : 1 ≤ n) : ContinuousOn (fun x => fderiv 𝕜 f x) s := ((contDiffOn_succ_iff_fderiv_of_isOpen hs).1 (h.of_le hn)).2.continuousOn +/-- The following lemma will be removed when the definition of `C^ω` will be corrected. For now, +it is only there as a convenient shortcut. -/ +theorem contDiffOn_infty_iff_contDiffOn_omega : + ContDiffOn 𝕜 ∞ f s ↔ ContDiffOn 𝕜 ω f s := by + have A (m : ℕ) : m ≤ ∞ := mod_cast le_top + simp [ContDiffOn, ContDiffWithinAt, hasFTaylorSeriesUpTo_top_iff, A] + /-! ### Smooth functions at a point -/ variable (𝕜) @@ -711,7 +733,7 @@ variable (𝕜) /-- A function is continuously differentiable up to `n` at a point `x` if, for any integer `k ≤ n`, there is a neighborhood of `x` where `f` admits derivatives up to order `n`, which are continuous. -/ -def ContDiffAt (n : ℕ∞) (f : E → F) (x : E) : Prop := +def ContDiffAt (n : WithTop ℕ∞) (f : E → F) (x : E) : Prop := ContDiffWithinAt 𝕜 n f univ x variable {𝕜} @@ -795,7 +817,7 @@ variable (𝕜) order `n`, which are continuous. Contrary to the case of definitions in domains (where derivatives might not be unique) we do not need to localize the definition in space or time. -/ -def ContDiff (n : ℕ∞) (f : E → F) : Prop := +def ContDiff (n : WithTop ℕ∞) (f : E → F) : Prop := ∃ p : E → FormalMultilinearSeries 𝕜 E F, HasFTaylorSeriesUpTo n f p variable {𝕜} @@ -824,10 +846,16 @@ theorem ContDiff.contDiffAt (h : ContDiff 𝕜 n f) : ContDiffAt 𝕜 n f x := theorem ContDiff.contDiffWithinAt (h : ContDiff 𝕜 n f) : ContDiffWithinAt 𝕜 n f s x := h.contDiffAt.contDiffWithinAt +/-- The following lemma will be removed when the definition of `C^ω` will be corrected. For now, +it is only there as a convenient shortcut. -/ +theorem contDiff_infty_iff_contDiff_omega : + ContDiff 𝕜 ∞ f ↔ ContDiff 𝕜 ω f := by + simp [ContDiff, hasFTaylorSeriesUpTo_top_iff] + theorem contDiff_top : ContDiff 𝕜 ∞ f ↔ ∀ n : ℕ, ContDiff 𝕜 n f := by simp [contDiffOn_univ.symm, contDiffOn_top] -theorem contDiff_all_iff_nat : (∀ n, ContDiff 𝕜 n f) ↔ ∀ n : ℕ, ContDiff 𝕜 n f := by +theorem contDiff_all_iff_nat : (∀ (n : ℕ∞), ContDiff 𝕜 n f) ↔ ∀ n : ℕ, ContDiff 𝕜 n f := by simp only [← contDiffOn_univ, contDiffOn_all_iff_nat] theorem ContDiff.contDiffOn (h : ContDiff 𝕜 n f) : ContDiffOn 𝕜 n f s := @@ -844,8 +872,8 @@ theorem contDiffAt_zero : ContDiffAt 𝕜 0 f x ↔ ∃ u ∈ 𝓝 x, Continuous theorem contDiffAt_one_iff : ContDiffAt 𝕜 1 f x ↔ ∃ f' : E → E →L[𝕜] F, ∃ u ∈ 𝓝 x, ContinuousOn f' u ∧ ∀ x ∈ u, HasFDerivAt f (f' x) x := by - rw [show (1 : ℕ∞) = (0 : ℕ) + 1 from rfl, contDiffAt_succ_iff_hasFDerivAt] - simp_rw [show ((0 : ℕ) : ℕ∞) = 0 from rfl, contDiffAt_zero, + rw [show (1 : WithTop ℕ∞) = (0 : ℕ) + 1 from rfl, contDiffAt_succ_iff_hasFDerivAt] + simp_rw [show ((0 : ℕ) : WithTop ℕ∞) = 0 from rfl, contDiffAt_zero, exists_mem_and_iff antitone_bforall antitone_continuousOn, and_comm] theorem ContDiff.of_le (h : ContDiff 𝕜 n f) (hmn : m ≤ n) : ContDiff 𝕜 m f := @@ -864,7 +892,8 @@ theorem ContDiff.continuous (h : ContDiff 𝕜 n f) : Continuous f := theorem ContDiff.differentiable (h : ContDiff 𝕜 n f) (hn : 1 ≤ n) : Differentiable 𝕜 f := differentiableOn_univ.1 <| (contDiffOn_univ.2 h).differentiableOn hn -theorem contDiff_iff_forall_nat_le : ContDiff 𝕜 n f ↔ ∀ m : ℕ, ↑m ≤ n → ContDiff 𝕜 m f := by +theorem contDiff_iff_forall_nat_le {n : ℕ∞} : + ContDiff 𝕜 n f ↔ ∀ m : ℕ, ↑m ≤ n → ContDiff 𝕜 m f := by simp_rw [← contDiffOn_univ]; exact contDiffOn_iff_forall_nat_le /-- A function is `C^(n+1)` iff it has a `C^n` derivative. -/ @@ -887,7 +916,7 @@ theorem contDiff_iff_ftaylorSeries : exact fun h => ContDiffOn.ftaylorSeriesWithin h uniqueDiffOn_univ · intro h; exact ⟨ftaylorSeries 𝕜 f, h⟩ -theorem contDiff_iff_continuous_differentiable : +theorem contDiff_iff_continuous_differentiable {n : ℕ∞} : ContDiff 𝕜 n f ↔ (∀ m : ℕ, m ≤ n → Continuous fun x => iteratedFDeriv 𝕜 m f x) ∧ ∀ m : ℕ, m < n → Differentiable 𝕜 fun x => iteratedFDeriv 𝕜 m f x := by @@ -897,14 +926,15 @@ theorem contDiff_iff_continuous_differentiable : /-- If `f` is `C^n` then its `m`-times iterated derivative is continuous for `m ≤ n`. -/ theorem ContDiff.continuous_iteratedFDeriv {m : ℕ} (hm : m ≤ n) (hf : ContDiff 𝕜 n f) : Continuous fun x => iteratedFDeriv 𝕜 m f x := - (contDiff_iff_continuous_differentiable.mp hf).1 m hm + (contDiff_iff_continuous_differentiable.mp (hf.of_le hm)).1 m le_rfl /-- If `f` is `C^n` then its `m`-times iterated derivative is differentiable for `m < n`. -/ theorem ContDiff.differentiable_iteratedFDeriv {m : ℕ} (hm : m < n) (hf : ContDiff 𝕜 n f) : Differentiable 𝕜 fun x => iteratedFDeriv 𝕜 m f x := - (contDiff_iff_continuous_differentiable.mp hf).2 m hm + (contDiff_iff_continuous_differentiable.mp + (hf.of_le (ENat.add_one_natCast_le_withTop_of_lt hm))).2 m (mod_cast lt_add_one m) -theorem contDiff_of_differentiable_iteratedFDeriv +theorem contDiff_of_differentiable_iteratedFDeriv {n : ℕ∞} (h : ∀ m : ℕ, m ≤ n → Differentiable 𝕜 (iteratedFDeriv 𝕜 m f)) : ContDiff 𝕜 n f := contDiff_iff_continuous_differentiable.2 ⟨fun m hm => (h m hm).continuous, fun m hm => h m (le_of_lt hm)⟩ diff --git a/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean b/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean index 555d81bb107e4..c399175ddd00d 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean @@ -20,6 +20,7 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {D : Type uD} [NormedAddCommGroup D] [NormedSpace 𝕜 D] {E : Type uE} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {F : Type uF} [NormedAddCommGroup F] [NormedSpace 𝕜 F] + {n : WithTop ℕ∞} {f : D → E} {s : Set D} /-! ### Finite dimensional results -/ @@ -30,7 +31,7 @@ open Function Module variable [CompleteSpace 𝕜] /-- A family of continuous linear maps is `C^n` on `s` if all its applications are. -/ -theorem contDiffOn_clm_apply {n : ℕ∞} {f : D → E →L[𝕜] F} {s : Set D} [FiniteDimensional 𝕜 E] : +theorem contDiffOn_clm_apply {f : D → E →L[𝕜] F} {s : Set D} [FiniteDimensional 𝕜 E] : ContDiffOn 𝕜 n f s ↔ ∀ y, ContDiffOn 𝕜 n (fun x => f x y) s := by refine ⟨fun h y => h.clm_apply contDiffOn_const, fun h => ?_⟩ let d := finrank 𝕜 E @@ -40,7 +41,7 @@ theorem contDiffOn_clm_apply {n : ℕ∞} {f : D → E →L[𝕜] F} {s : Set D} rw [← id_comp f, ← e₂.symm_comp_self] exact e₂.symm.contDiff.comp_contDiffOn (contDiffOn_pi.mpr fun i => h _) -theorem contDiff_clm_apply_iff {n : ℕ∞} {f : D → E →L[𝕜] F} [FiniteDimensional 𝕜 E] : +theorem contDiff_clm_apply_iff {f : D → E →L[𝕜] F} [FiniteDimensional 𝕜 E] : ContDiff 𝕜 n f ↔ ∀ y, ContDiff 𝕜 n fun x => f x y := by simp_rw [← contDiffOn_univ, contDiffOn_clm_apply] @@ -48,7 +49,7 @@ theorem contDiff_clm_apply_iff {n : ℕ∞} {f : D → E →L[𝕜] F} [FiniteDi When you do induction on `n`, this gives a useful characterization of a function being `C^(n+1)`, assuming you have already computed the derivative. The advantage of this version over `contDiff_succ_iff_fderiv` is that both occurrences of `ContDiff` are for functions with the same -domain and codomain (`E` and `F`). This is not the case for `contDiff_succ_iff_fderiv`, which +domain and codomain (`D` and `E`). This is not the case for `contDiff_succ_iff_fderiv`, which often requires an inconvenient need to generalize `F`, which results in universe issues (see the discussion in the section of `ContDiff.comp`). diff --git a/Mathlib/Analysis/Calculus/ContDiff/RCLike.lean b/Mathlib/Analysis/Calculus/ContDiff/RCLike.lean index 5f8fdba324bcb..e87af872c43a2 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/RCLike.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/RCLike.lean @@ -24,8 +24,8 @@ section Real its extension fields such as `ℂ`). -/ -variable {n : ℕ∞} {𝕂 : Type*} [RCLike 𝕂] {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕂 E'] - {F' : Type*} [NormedAddCommGroup F'] [NormedSpace 𝕂 F'] +variable {n : WithTop ℕ∞} {𝕂 : Type*} [RCLike 𝕂] {E' : Type*} [NormedAddCommGroup E'] + [NormedSpace 𝕂 E'] {F' : Type*} [NormedAddCommGroup F'] [NormedSpace 𝕂 F'] /-- If a function has a Taylor series at order at least 1, then at points in the interior of the domain of definition, the term of order 1 of this series is a strict derivative of `f`. -/ @@ -41,7 +41,7 @@ us as `f'`, then `f'` is also a strict derivative. -/ theorem ContDiffAt.hasStrictFDerivAt' {f : E' → F'} {f' : E' →L[𝕂] F'} {x : E'} (hf : ContDiffAt 𝕂 n f x) (hf' : HasFDerivAt f f' x) (hn : 1 ≤ n) : HasStrictFDerivAt f f' x := by - rcases hf 1 hn with ⟨u, H, p, hp⟩ + rcases hf.of_le hn 1 le_rfl with ⟨u, H, p, hp⟩ simp only [nhdsWithin_univ, mem_univ, insert_eq_of_mem] at H have := hp.hasStrictFDerivAt le_rfl H rwa [hf'.unique this.hasFDerivAt] @@ -135,7 +135,7 @@ lemma ContDiff.locallyLipschitz {f : E' → F'} (hf : ContDiff 𝕂 1 f) : Local use K, t /-- A `C^1` function with compact support is Lipschitz. -/ -theorem ContDiff.lipschitzWith_of_hasCompactSupport {f : E' → F'} {n : ℕ∞} +theorem ContDiff.lipschitzWith_of_hasCompactSupport {f : E' → F'} (hf : HasCompactSupport f) (h'f : ContDiff 𝕂 n f) (hn : 1 ≤ n) : ∃ C, LipschitzWith C f := by obtain ⟨C, hC⟩ := (hf.fderiv 𝕂).exists_bound_of_continuous (h'f.continuous_fderiv hn) diff --git a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean index 055c8cb25720e..9141ffab19f41 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean @@ -63,7 +63,7 @@ differentiability at points in a neighborhood of `s`. Therefore, the theorem tha open Filter Asymptotics Set -open scoped ENNReal Topology +open scoped ENNReal Topology ContDiff universe u v @@ -261,7 +261,7 @@ by the sequence of its derivatives. Note that, if the function were just analyti 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) : + (n : WithTop ℕ∞) (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 @@ -270,25 +270,27 @@ lemma AnalyticOnNhd.hasFTaylorSeriesUpToOn [CompleteSpace F] exact (h.iteratedFDeriv m x hx).differentiableAt /-- An analytic function is infinitely differentiable. -/ -protected theorem AnalyticOnNhd.contDiffOn [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f s) {n : ℕ∞} : - ContDiffOn 𝕜 n f s := +protected theorem AnalyticOnNhd.contDiffOn [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f s) + {n : WithTop ℕ∞} : ContDiffOn 𝕜 n f s := by + suffices ContDiffOn 𝕜 ω f s from this.of_le le_top + rw [← contDiffOn_infty_iff_contDiffOn_omega] let t := { x | AnalyticAt 𝕜 f x } - suffices ContDiffOn 𝕜 n f t from this.mono h + suffices ContDiffOn 𝕜 ∞ f t from this.mono h have H : AnalyticOnNhd 𝕜 f t := fun _x hx ↦ hx have t_open : IsOpen t := isOpen_analyticAt 𝕜 f - contDiffOn_of_continuousOn_differentiableOn + exact contDiffOn_of_continuousOn_differentiableOn (fun m _ ↦ (H.iteratedFDeriv m).continuousOn.congr fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) (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 : ℕ∞} : +theorem AnalyticOnNhd.contDiff [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f univ) {n : WithTop ℕ∞} : ContDiff 𝕜 n f := by rw [← contDiffOn_univ] exact h.contDiffOn -theorem AnalyticAt.contDiffAt [CompleteSpace F] (h : AnalyticAt 𝕜 f x) {n : ℕ∞} : +theorem AnalyticAt.contDiffAt [CompleteSpace F] (h : AnalyticAt 𝕜 f x) {n : WithTop ℕ∞} : ContDiffAt 𝕜 n f x := by obtain ⟨s, hs, hf⟩ := h.exists_mem_nhds_analyticOnNhd exact hf.contDiffOn.contDiffAt hs @@ -306,7 +308,7 @@ protected lemma AnalyticOn.contDiffOn [CompleteSpace F] {f : E → F} {s : Set E alias AnalyticWithinOn.contDiffOn := AnalyticOn.contDiffOn lemma AnalyticWithinAt.exists_hasFTaylorSeriesUpToOn [CompleteSpace F] - (n : ℕ∞) (h : AnalyticWithinAt 𝕜 f s x) : + (n : WithTop ℕ∞) (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⟩ @@ -545,19 +547,21 @@ theorem CPolynomialOn.iteratedFDeriv (h : CPolynomialOn 𝕜 f s) (n : ℕ) : simp /-- A polynomial function is infinitely differentiable. -/ -theorem CPolynomialOn.contDiffOn (h : CPolynomialOn 𝕜 f s) {n : ℕ∞} : - ContDiffOn 𝕜 n f s := +theorem CPolynomialOn.contDiffOn (h : CPolynomialOn 𝕜 f s) {n : WithTop ℕ∞} : + ContDiffOn 𝕜 n f s := by + suffices ContDiffOn 𝕜 ω f s from this.of_le le_top let t := { x | CPolynomialAt 𝕜 f x } - suffices ContDiffOn 𝕜 n f t from this.mono h + suffices ContDiffOn 𝕜 ω f t from this.mono h + rw [← contDiffOn_infty_iff_contDiffOn_omega] have H : CPolynomialOn 𝕜 f t := fun _x hx ↦ hx have t_open : IsOpen t := isOpen_cPolynomialAt 𝕜 f - contDiffOn_of_continuousOn_differentiableOn + exact contDiffOn_of_continuousOn_differentiableOn (fun m _ ↦ (H.iteratedFDeriv m).continuousOn.congr fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) (fun m _ ↦ (H.iteratedFDeriv m).analyticOnNhd.differentiableOn.congr fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) -theorem CPolynomialAt.contDiffAt (h : CPolynomialAt 𝕜 f x) {n : ℕ∞} : +theorem CPolynomialAt.contDiffAt (h : CPolynomialAt 𝕜 f x) {n : WithTop ℕ∞} : ContDiffAt 𝕜 n f x := let ⟨_, hs, hf⟩ := h.exists_mem_nhds_cPolynomialOn hf.contDiffOn.contDiffAt hs @@ -595,7 +599,7 @@ theorem changeOriginSeries_support {k l : ℕ} (h : k + l ≠ Fintype.card ι) : simp_rw [FormalMultilinearSeries.changeOriginSeriesTerm, toFormalMultilinearSeries, dif_neg h.symm, LinearIsometryEquiv.map_zero] -variable {n : ℕ∞} (x : ∀ i, E i) +variable {n : WithTop ℕ∞} (x : ∀ i, E i) open Finset in theorem changeOrigin_toFormalMultilinearSeries [DecidableEq ι] : diff --git a/Mathlib/Analysis/Calculus/FDeriv/Norm.lean b/Mathlib/Analysis/Calculus/FDeriv/Norm.lean index 4587f2bbf6dfc..fabdde33a4bfd 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Norm.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Norm.lean @@ -38,7 +38,7 @@ differentiability, norm open ContinuousLinearMap Filter NNReal Real Set variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] -variable {n : ℕ∞} {f : E →L[ℝ] ℝ} {x : E} {t : ℝ} +variable {n : WithTop ℕ∞} {f : E →L[ℝ] ℝ} {x : E} {t : ℝ} variable (E) in theorem not_differentiableAt_norm_zero [Nontrivial E] : @@ -72,15 +72,15 @@ theorem contDiffAt_norm_smul_iff (ht : t ≠ 0) : theorem ContDiffAt.contDiffAt_norm_of_smul (h : ContDiffAt ℝ n (‖·‖) (t • x)) : ContDiffAt ℝ n (‖·‖) x := by - obtain rfl | hn : n = 0 ∨ 1 ≤ n := by - rw [← ENat.lt_one_iff_eq_zero] - exact lt_or_le .. - · rw [contDiffAt_zero] + rcases eq_bot_or_bot_lt n with rfl | hn + · apply contDiffAt_zero.2 exact ⟨univ, univ_mem, continuous_norm.continuousOn⟩ + replace hn : 1 ≤ n := ENat.add_one_natCast_le_withTop_of_lt hn obtain rfl | ht := eq_or_ne t 0 · by_cases hE : Nontrivial E · rw [zero_smul] at h - exact (mt (ContDiffAt.differentiableAt · hn)) (not_differentiableAt_norm_zero E) h |>.elim + exact (mt (ContDiffAt.differentiableAt · (mod_cast hn))) + (not_differentiableAt_norm_zero E) h |>.elim · rw [not_nontrivial_iff_subsingleton] at hE rw [eq_const_of_subsingleton (‖·‖) 0] exact contDiffAt_const diff --git a/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean b/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean index a893dcc67f3d8..b4f9b4db76e68 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean @@ -444,7 +444,7 @@ theorem second_derivative_symmetric {f' : E → E →L[𝕜] F} {f'' : E →L[ second_derivative_symmetric_of_eventually (Filter.Eventually.of_forall hf) hx v w /-- If a function is `C^2` at a point, then its second derivative there is symmetric. -/ -theorem ContDiffAt.isSymmSndFDerivAt {n : ℕ∞} (hf : ContDiffAt 𝕜 n f x) (hn : 2 ≤ n) : +theorem ContDiffAt.isSymmSndFDerivAt {n : WithTop ℕ∞} (hf : ContDiffAt 𝕜 n f x) (hn : 2 ≤ n) : IsSymmSndFDerivAt 𝕜 f x := by intro v w apply second_derivative_symmetric_of_eventually (f := f) (f' := fderiv 𝕜 f) (x := x) @@ -462,7 +462,7 @@ theorem ContDiffAt.isSymmSndFDerivAt {n : ℕ∞} (hf : ContDiffAt 𝕜 n f x) ( /-- If a function is `C^2` within a set at a point, and accumulated by points in the interior of the set, then its second derivative there is symmetric. -/ -theorem ContDiffWithinAt.isSymmSndFDerivWithinAt {n : ℕ∞} (hf : ContDiffWithinAt 𝕜 n f s x) +theorem ContDiffWithinAt.isSymmSndFDerivWithinAt {n : WithTop ℕ∞} (hf : ContDiffWithinAt 𝕜 n f s x) (hn : 2 ≤ n) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ closure (interior s)) (h'x : x ∈ s) : IsSymmSndFDerivWithinAt 𝕜 f s x := by /- We argue that, at interior points, the second derivative is symmetric, and moreover by diff --git a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ContDiff.lean b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ContDiff.lean index 55e18e301150a..49b5c3a3c929c 100644 --- a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ContDiff.lean +++ b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ContDiff.lean @@ -20,46 +20,46 @@ namespace ContDiffAt variable {𝕂 : Type*} [RCLike 𝕂] variable {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕂 E] variable {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕂 F] -variable [CompleteSpace E] (f : E → F) {f' : E ≃L[𝕂] F} {a : E} +variable [CompleteSpace E] (f : E → F) {f' : E ≃L[𝕂] F} {a : E} {n : WithTop ℕ∞} /-- Given a `ContDiff` function over `𝕂` (which is `ℝ` or `ℂ`) with an invertible derivative at `a`, returns a `PartialHomeomorph` with `to_fun = f` and `a ∈ source`. -/ -def toPartialHomeomorph {n : ℕ∞} (hf : ContDiffAt 𝕂 n f a) (hf' : HasFDerivAt f (f' : E →L[𝕂] F) a) +def toPartialHomeomorph (hf : ContDiffAt 𝕂 n f a) (hf' : HasFDerivAt f (f' : E →L[𝕂] F) a) (hn : 1 ≤ n) : PartialHomeomorph E F := (hf.hasStrictFDerivAt' hf' hn).toPartialHomeomorph f variable {f} @[simp] -theorem toPartialHomeomorph_coe {n : ℕ∞} (hf : ContDiffAt 𝕂 n f a) +theorem toPartialHomeomorph_coe (hf : ContDiffAt 𝕂 n f a) (hf' : HasFDerivAt f (f' : E →L[𝕂] F) a) (hn : 1 ≤ n) : (hf.toPartialHomeomorph f hf' hn : E → F) = f := rfl -theorem mem_toPartialHomeomorph_source {n : ℕ∞} (hf : ContDiffAt 𝕂 n f a) +theorem mem_toPartialHomeomorph_source (hf : ContDiffAt 𝕂 n f a) (hf' : HasFDerivAt f (f' : E →L[𝕂] F) a) (hn : 1 ≤ n) : a ∈ (hf.toPartialHomeomorph f hf' hn).source := (hf.hasStrictFDerivAt' hf' hn).mem_toPartialHomeomorph_source -theorem image_mem_toPartialHomeomorph_target {n : ℕ∞} (hf : ContDiffAt 𝕂 n f a) +theorem image_mem_toPartialHomeomorph_target (hf : ContDiffAt 𝕂 n f a) (hf' : HasFDerivAt f (f' : E →L[𝕂] F) a) (hn : 1 ≤ n) : f a ∈ (hf.toPartialHomeomorph f hf' hn).target := (hf.hasStrictFDerivAt' hf' hn).image_mem_toPartialHomeomorph_target /-- Given a `ContDiff` function over `𝕂` (which is `ℝ` or `ℂ`) with an invertible derivative at `a`, returns a function that is locally inverse to `f`. -/ -def localInverse {n : ℕ∞} (hf : ContDiffAt 𝕂 n f a) (hf' : HasFDerivAt f (f' : E →L[𝕂] F) a) +def localInverse (hf : ContDiffAt 𝕂 n f a) (hf' : HasFDerivAt f (f' : E →L[𝕂] F) a) (hn : 1 ≤ n) : F → E := (hf.hasStrictFDerivAt' hf' hn).localInverse f f' a -theorem localInverse_apply_image {n : ℕ∞} (hf : ContDiffAt 𝕂 n f a) +theorem localInverse_apply_image (hf : ContDiffAt 𝕂 n f a) (hf' : HasFDerivAt f (f' : E →L[𝕂] F) a) (hn : 1 ≤ n) : hf.localInverse hf' hn (f a) = a := (hf.hasStrictFDerivAt' hf' hn).localInverse_apply_image /-- Given a `ContDiff` function over `𝕂` (which is `ℝ` or `ℂ`) with an invertible derivative at `a`, the inverse function (produced by `ContDiff.toPartialHomeomorph`) is also `ContDiff`. -/ -theorem to_localInverse {n : ℕ∞} (hf : ContDiffAt 𝕂 n f a) +theorem to_localInverse (hf : ContDiffAt 𝕂 n f a) (hf' : HasFDerivAt f (f' : E →L[𝕂] F) a) (hn : 1 ≤ n) : ContDiffAt 𝕂 n (hf.localInverse hf' hn) (f a) := by have := hf.localInverse_apply_image hf' hn diff --git a/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean b/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean index 9fcaa3d13a0cc..858642c83caa8 100644 --- a/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean +++ b/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean @@ -131,12 +131,13 @@ theorem contDiffOn_of_differentiableOn_deriv {n : ℕ∞} /-- On a set with unique derivatives, a `C^n` function has derivatives up to `n` which are continuous. -/ -theorem ContDiffOn.continuousOn_iteratedDerivWithin {n : ℕ∞} {m : ℕ} (h : ContDiffOn 𝕜 n f s) +theorem ContDiffOn.continuousOn_iteratedDerivWithin + {n : WithTop ℕ∞} {m : ℕ} (h : ContDiffOn 𝕜 n f s) (hmn : (m : ℕ∞) ≤ n) (hs : UniqueDiffOn 𝕜 s) : ContinuousOn (iteratedDerivWithin m f s) s := by simpa only [iteratedDerivWithin_eq_equiv_comp, LinearIsometryEquiv.comp_continuousOn_iff] using h.continuousOn_iteratedFDerivWithin hmn hs -theorem ContDiffWithinAt.differentiableWithinAt_iteratedDerivWithin {n : ℕ∞} {m : ℕ} +theorem ContDiffWithinAt.differentiableWithinAt_iteratedDerivWithin {n : WithTop ℕ∞} {m : ℕ} (h : ContDiffWithinAt 𝕜 n f s x) (hmn : (m : ℕ∞) < n) (hs : UniqueDiffOn 𝕜 (insert x s)) : DifferentiableWithinAt 𝕜 (iteratedDerivWithin m f s) s x := by simpa only [iteratedDerivWithin_eq_equiv_comp, @@ -145,8 +146,8 @@ theorem ContDiffWithinAt.differentiableWithinAt_iteratedDerivWithin {n : ℕ∞} /-- On a set with unique derivatives, a `C^n` function has derivatives less than `n` which are differentiable. -/ -theorem ContDiffOn.differentiableOn_iteratedDerivWithin {n : ℕ∞} {m : ℕ} (h : ContDiffOn 𝕜 n f s) - (hmn : (m : ℕ∞) < n) (hs : UniqueDiffOn 𝕜 s) : +theorem ContDiffOn.differentiableOn_iteratedDerivWithin {n : WithTop ℕ∞} {m : ℕ} + (h : ContDiffOn 𝕜 n f s) (hmn : m < n) (hs : UniqueDiffOn 𝕜 s) : DifferentiableOn 𝕜 (iteratedDerivWithin m f s) s := fun x hx => (h x hx).differentiableWithinAt_iteratedDerivWithin hmn <| by rwa [insert_eq_of_mem hx] @@ -238,13 +239,14 @@ theorem contDiff_of_differentiable_iteratedDeriv {n : ℕ∞} (h : ∀ m : ℕ, (m : ℕ∞) ≤ n → Differentiable 𝕜 (iteratedDeriv m f)) : ContDiff 𝕜 n f := contDiff_iff_iteratedDeriv.2 ⟨fun m hm => (h m hm).continuous, fun m hm => h m (le_of_lt hm)⟩ -theorem ContDiff.continuous_iteratedDeriv {n : ℕ∞} (m : ℕ) (h : ContDiff 𝕜 n f) - (hmn : (m : ℕ∞) ≤ n) : Continuous (iteratedDeriv m f) := - (contDiff_iff_iteratedDeriv.1 h).1 m hmn +theorem ContDiff.continuous_iteratedDeriv {n : WithTop ℕ∞} (m : ℕ) (h : ContDiff 𝕜 n f) + (hmn : m ≤ n) : Continuous (iteratedDeriv m f) := + (contDiff_iff_iteratedDeriv.1 (h.of_le hmn)).1 m le_rfl -theorem ContDiff.differentiable_iteratedDeriv {n : ℕ∞} (m : ℕ) (h : ContDiff 𝕜 n f) - (hmn : (m : ℕ∞) < n) : Differentiable 𝕜 (iteratedDeriv m f) := - (contDiff_iff_iteratedDeriv.1 h).2 m hmn +theorem ContDiff.differentiable_iteratedDeriv {n : WithTop ℕ∞} (m : ℕ) (h : ContDiff 𝕜 n f) + (hmn : m < n) : Differentiable 𝕜 (iteratedDeriv m f) := + (contDiff_iff_iteratedDeriv.1 (h.of_le (ENat.add_one_natCast_le_withTop_of_lt hmn))).2 m + (mod_cast (lt_add_one m)) /-- The `n+1`-th iterated derivative can be obtained by differentiating the `n`-th iterated derivative. -/ diff --git a/Mathlib/Analysis/Calculus/Rademacher.lean b/Mathlib/Analysis/Calculus/Rademacher.lean index 28ad0559741a1..1c836a3a36f92 100644 --- a/Mathlib/Analysis/Calculus/Rademacher.lean +++ b/Mathlib/Analysis/Calculus/Rademacher.lean @@ -221,9 +221,9 @@ theorem ae_lineDeriv_sum_eq suffices S2 : ∫ x, (∑ i ∈ s, a i * fderiv ℝ g x (v i)) * f x ∂μ = ∑ i ∈ s, a i * ∫ x, fderiv ℝ g x (v i) * f x ∂μ by obtain ⟨D, g_lip⟩ : ∃ D, LipschitzWith D g := - ContDiff.lipschitzWith_of_hasCompactSupport g_comp g_smooth le_top + ContDiff.lipschitzWith_of_hasCompactSupport g_comp g_smooth (mod_cast le_top) simp_rw [integral_lineDeriv_mul_eq hf g_lip g_comp] - simp_rw [(g_smooth.differentiable le_top).differentiableAt.lineDeriv_eq_fderiv] + simp_rw [(g_smooth.differentiable (mod_cast le_top)).differentiableAt.lineDeriv_eq_fderiv] simp only [map_neg, _root_.map_sum, _root_.map_smul, smul_eq_mul, neg_mul] simp only [integral_neg, mul_neg, Finset.sum_neg_distrib, neg_inj] exact S2 @@ -233,7 +233,8 @@ theorem ae_lineDeriv_sum_eq let L : (E →L[ℝ] ℝ) → ℝ := fun f ↦ f (v i) change Integrable (fun x ↦ a i * ((L ∘ (fderiv ℝ g)) x * f x)) μ refine (Continuous.integrable_of_hasCompactSupport ?_ ?_).const_mul _ - · exact ((g_smooth.continuous_fderiv le_top).clm_apply continuous_const).mul hf.continuous + · exact ((g_smooth.continuous_fderiv (mod_cast le_top)).clm_apply continuous_const).mul + hf.continuous · exact ((g_comp.fderiv ℝ).comp_left rfl).mul_right /-! diff --git a/Mathlib/Analysis/Calculus/SmoothSeries.lean b/Mathlib/Analysis/Calculus/SmoothSeries.lean index 58d0f39d86c54..073a4b2192939 100644 --- a/Mathlib/Analysis/Calculus/SmoothSeries.lean +++ b/Mathlib/Analysis/Calculus/SmoothSeries.lean @@ -198,7 +198,8 @@ theorem iteratedFDeriv_tsum (hf : ∀ i, ContDiff 𝕜 N (f i)) have A : Summable fun n => iteratedFDeriv 𝕜 k (f n) 0 := .of_norm_bounded (v k) (hv k h'k.le) fun n => h'f k n 0 h'k.le simp_rw [iteratedFDeriv_succ_eq_comp_left, IH h'k.le] - rw [fderiv_tsum (hv _ hk) (fun n => (hf n).differentiable_iteratedFDeriv h'k) _ A] + rw [fderiv_tsum (hv _ hk) (fun n => (hf n).differentiable_iteratedFDeriv + (mod_cast h'k)) _ A] · ext1 x exact (continuousMultilinearCurryLeftEquiv 𝕜 (fun _ : Fin (k + 1) => E) F).symm.toContinuousLinearEquiv.map_tsum @@ -219,7 +220,7 @@ theorem iteratedFDeriv_tsum_apply (hf : ∀ i, ContDiff 𝕜 N (f i)) class `C^N`, and moreover there is a uniform summable upper bound on the `k`-th derivative for each `k ≤ N`. Then the series is also `C^N`. -/ theorem contDiff_tsum (hf : ∀ i, ContDiff 𝕜 N (f i)) (hv : ∀ k : ℕ, (k : ℕ∞) ≤ N → Summable (v k)) - (h'f : ∀ (k : ℕ) (i : α) (x : E), (k : ℕ∞) ≤ N → ‖iteratedFDeriv 𝕜 k (f i) x‖ ≤ v k i) : + (h'f : ∀ (k : ℕ) (i : α) (x : E), k ≤ N → ‖iteratedFDeriv 𝕜 k (f i) x‖ ≤ v k i) : ContDiff 𝕜 N fun x => ∑' i, f i x := by rw [contDiff_iff_continuous_differentiable] constructor @@ -227,16 +228,16 @@ theorem contDiff_tsum (hf : ∀ i, ContDiff 𝕜 N (f i)) (hv : ∀ k : ℕ, (k rw [iteratedFDeriv_tsum hf hv h'f hm] refine continuous_tsum ?_ (hv m hm) ?_ · intro i - exact ContDiff.continuous_iteratedFDeriv hm (hf i) + exact ContDiff.continuous_iteratedFDeriv (mod_cast hm) (hf i) · intro n x exact h'f _ _ _ hm · intro m hm have h'm : ((m + 1 : ℕ) : ℕ∞) ≤ N := by simpa only [ENat.coe_add, ENat.coe_one] using Order.add_one_le_of_lt hm rw [iteratedFDeriv_tsum hf hv h'f hm.le] - have A : - ∀ n x, HasFDerivAt (iteratedFDeriv 𝕜 m (f n)) (fderiv 𝕜 (iteratedFDeriv 𝕜 m (f n)) x) x := - fun n x => (ContDiff.differentiable_iteratedFDeriv hm (hf n)).differentiableAt.hasFDerivAt + have A n x : HasFDerivAt (iteratedFDeriv 𝕜 m (f n)) (fderiv 𝕜 (iteratedFDeriv 𝕜 m (f n)) x) x := + (ContDiff.differentiable_iteratedFDeriv (mod_cast hm) + (hf n)).differentiableAt.hasFDerivAt refine differentiable_tsum (hv _ h'm) A fun n x => ?_ rw [fderiv_iteratedFDeriv, comp_apply, LinearIsometryEquiv.norm_map] exact h'f _ _ _ h'm @@ -245,11 +246,9 @@ theorem contDiff_tsum (hf : ∀ i, ContDiff 𝕜 N (f i)) (hv : ∀ k : ℕ, (k class `C^N`, and moreover there is a uniform summable upper bound on the `k`-th derivative for each `k ≤ N` (except maybe for finitely many `i`s). Then the series is also `C^N`. -/ theorem contDiff_tsum_of_eventually (hf : ∀ i, ContDiff 𝕜 N (f i)) - (hv : ∀ k : ℕ, (k : ℕ∞) ≤ N → Summable (v k)) - (h'f : - ∀ k : ℕ, - (k : ℕ∞) ≤ N → - ∀ᶠ i in (Filter.cofinite : Filter α), ∀ x : E, ‖iteratedFDeriv 𝕜 k (f i) x‖ ≤ v k i) : + (hv : ∀ k : ℕ, k ≤ N → Summable (v k)) + (h'f : ∀ k : ℕ, k ≤ N → + ∀ᶠ i in (Filter.cofinite : Filter α), ∀ x : E, ‖iteratedFDeriv 𝕜 k (f i) x‖ ≤ v k i) : ContDiff 𝕜 N fun x => ∑' i, f i x := by classical refine contDiff_iff_forall_nat_le.2 fun m hm => ?_ @@ -274,10 +273,10 @@ theorem contDiff_tsum_of_eventually (hf : ∀ i, ContDiff 𝕜 N (f i)) filter_upwards [h'f 0 (zero_le _)] with i hi simpa only [norm_iteratedFDeriv_zero] using hi x rw [this] - apply (ContDiff.sum fun i _ => (hf i).of_le hm).add + apply (ContDiff.sum fun i _ => (hf i).of_le (mod_cast hm)).add have h'u : ∀ k : ℕ, (k : ℕ∞) ≤ m → Summable (v k ∘ ((↑) : { i // i ∉ T } → α)) := fun k hk => (hv k (hk.trans hm)).subtype _ - refine contDiff_tsum (fun i => (hf i).of_le hm) h'u ?_ + refine contDiff_tsum (fun i => (hf i).of_le (mod_cast hm)) h'u ?_ rintro k ⟨i, hi⟩ x hk simp only [t, T, Finite.mem_toFinset, mem_setOf_eq, Finset.mem_range, not_forall, not_le, exists_prop, not_exists, not_and, not_lt] at hi diff --git a/Mathlib/Analysis/Calculus/Taylor.lean b/Mathlib/Analysis/Calculus/Taylor.lean index 71293eb9d3058..8a56fcf0d80a4 100644 --- a/Mathlib/Analysis/Calculus/Taylor.lean +++ b/Mathlib/Analysis/Calculus/Taylor.lean @@ -119,7 +119,8 @@ theorem continuousOn_taylorWithinEval {f : ℝ → E} {x : ℝ} {n : ℕ} {s : S simp_rw [taylor_within_apply] refine continuousOn_finset_sum (Finset.range (n + 1)) fun i hi => ?_ refine (continuousOn_const.mul ((continuousOn_const.sub continuousOn_id).pow _)).smul ?_ - rw [contDiffOn_iff_continuousOn_differentiableOn_deriv hs] at hf + rw [show (n : WithTop ℕ∞) = (n : ℕ∞) by rfl, + contDiffOn_iff_continuousOn_differentiableOn_deriv hs] at hf cases' hf with hf_left specialize hf_left i simp only [Finset.mem_range] at hi @@ -179,7 +180,7 @@ theorem hasDerivWithinAt_taylorWithinEval {f : ℝ → E} {x y : ℝ} {n : ℕ} simp only [add_zero, Nat.factorial_succ, Nat.cast_mul, Nat.cast_add, Nat.cast_one] have coe_lt_succ : (k : WithTop ℕ) < k.succ := Nat.cast_lt.2 k.lt_succ_self have hdiff : DifferentiableOn ℝ (iteratedDerivWithin k f s) s' := - (hf.differentiableOn_iteratedDerivWithin coe_lt_succ hs_unique).mono h + (hf.differentiableOn_iteratedDerivWithin (mod_cast coe_lt_succ) hs_unique).mono h specialize hk hf.of_succ ((hdiff y hy).mono_of_mem_nhdsWithin hs') convert hk.add (hasDerivWithinAt_taylor_coeff_within hs'_unique (nhdsWithin_mono _ h self_mem_nhdsWithin) hf') using 1 @@ -305,7 +306,7 @@ theorem taylor_mean_remainder_bound {f : ℝ → E} {a b C x : ℝ} {n : ℕ} (h simp [hx] -- The nth iterated derivative is differentiable have hf' : DifferentiableOn ℝ (iteratedDerivWithin n f (Icc a b)) (Icc a b) := - hf.differentiableOn_iteratedDerivWithin (WithTop.coe_lt_coe.mpr n.lt_succ_self) + hf.differentiableOn_iteratedDerivWithin (mod_cast n.lt_succ_self) (uniqueDiffOn_Icc h) -- We can uniformly bound the derivative of the Taylor polynomial have h' : ∀ y ∈ Ico a x, diff --git a/Mathlib/Analysis/Calculus/VectorField.lean b/Mathlib/Analysis/Calculus/VectorField.lean index d4e54bc4485f9..ac4eb32fdc707 100644 --- a/Mathlib/Analysis/Calculus/VectorField.lean +++ b/Mathlib/Analysis/Calculus/VectorField.lean @@ -144,7 +144,7 @@ lemma lieBracket_swap : lieBracket 𝕜 V W x = - lieBracket 𝕜 W V x := by ext x; simp [lieBracket] lemma _root_.ContDiffWithinAt.lieBracketWithin_vectorField - {m n : ℕ∞} (hV : ContDiffWithinAt 𝕜 n V s x) + {m n : WithTop ℕ∞} (hV : ContDiffWithinAt 𝕜 n V s x) (hW : ContDiffWithinAt 𝕜 n W s x) (hs : UniqueDiffOn 𝕜 s) (hmn : m + 1 ≤ n) (hx : x ∈ s) : ContDiffWithinAt 𝕜 m (lieBracketWithin 𝕜 V W s) s x := by apply ContDiffWithinAt.sub @@ -153,19 +153,19 @@ lemma _root_.ContDiffWithinAt.lieBracketWithin_vectorField · exact ContDiffWithinAt.clm_apply (hV.fderivWithin_right hs hmn hx) (hW.of_le (le_trans le_self_add hmn)) -lemma _root_.ContDiffAt.lieBracket_vectorField {m n : ℕ∞} (hV : ContDiffAt 𝕜 n V x) +lemma _root_.ContDiffAt.lieBracket_vectorField {m n : WithTop ℕ∞} (hV : ContDiffAt 𝕜 n V x) (hW : ContDiffAt 𝕜 n W x) (hmn : m + 1 ≤ n) : ContDiffAt 𝕜 m (lieBracket 𝕜 V W) x := by rw [← contDiffWithinAt_univ] at hV hW ⊢ simp_rw [← lieBracketWithin_univ] exact hV.lieBracketWithin_vectorField hW uniqueDiffOn_univ hmn (mem_univ _) -lemma _root_.ContDiffOn.lieBracketWithin_vectorField {m n : ℕ∞} (hV : ContDiffOn 𝕜 n V s) +lemma _root_.ContDiffOn.lieBracketWithin_vectorField {m n : WithTop ℕ∞} (hV : ContDiffOn 𝕜 n V s) (hW : ContDiffOn 𝕜 n W s) (hs : UniqueDiffOn 𝕜 s) (hmn : m + 1 ≤ n) : ContDiffOn 𝕜 m (lieBracketWithin 𝕜 V W s) s := fun x hx ↦ (hV x hx).lieBracketWithin_vectorField (hW x hx) hs hmn hx -lemma _root_.ContDiff.lieBracket_vectorField {m n : ℕ∞} (hV : ContDiff 𝕜 n V) +lemma _root_.ContDiff.lieBracket_vectorField {m n : WithTop ℕ∞} (hV : ContDiff 𝕜 n V) (hW : ContDiff 𝕜 n W) (hmn : m + 1 ≤ n) : ContDiff 𝕜 m (lieBracket 𝕜 V W) := contDiff_iff_contDiffAt.2 (fun _ ↦ hV.contDiffAt.lieBracket_vectorField hW.contDiffAt hmn) diff --git a/Mathlib/Analysis/Complex/RealDeriv.lean b/Mathlib/Analysis/Complex/RealDeriv.lean index 3ce85a20eb6f5..1609fa5e357a4 100644 --- a/Mathlib/Analysis/Complex/RealDeriv.lean +++ b/Mathlib/Analysis/Complex/RealDeriv.lean @@ -77,14 +77,14 @@ theorem HasDerivAt.real_of_complex (h : HasDerivAt e e' z) : rw [ContinuousLinearMap.comp_apply, ContinuousLinearMap.comp_apply] simp -theorem ContDiffAt.real_of_complex {n : ℕ∞} (h : ContDiffAt ℂ n e z) : +theorem ContDiffAt.real_of_complex {n : WithTop ℕ∞} (h : ContDiffAt ℂ n e z) : ContDiffAt ℝ n (fun x : ℝ => (e x).re) z := by have A : ContDiffAt ℝ n ((↑) : ℝ → ℂ) z := ofRealCLM.contDiff.contDiffAt have B : ContDiffAt ℝ n e z := h.restrict_scalars ℝ have C : ContDiffAt ℝ n re (e z) := reCLM.contDiff.contDiffAt exact C.comp z (B.comp z A) -theorem ContDiff.real_of_complex {n : ℕ∞} (h : ContDiff ℂ n e) : +theorem ContDiff.real_of_complex {n : WithTop ℕ∞} (h : ContDiff ℂ n e) : ContDiff ℝ n fun x : ℝ => (e x).re := contDiff_iff_contDiffAt.2 fun _ => h.contDiffAt.real_of_complex diff --git a/Mathlib/Analysis/Convolution.lean b/Mathlib/Analysis/Convolution.lean index e3a80a3cd256e..f2e6fd1a550aa 100644 --- a/Mathlib/Analysis/Convolution.lean +++ b/Mathlib/Analysis/Convolution.lean @@ -1179,17 +1179,18 @@ theorem contDiffOn_convolution_right_with_param_aux {G : Type uP} {E' : Type uP} come from the same universe). -/ induction n using ENat.nat_induction generalizing g E' F with | h0 => - rw [contDiffOn_zero] at hg ⊢ + rw [WithTop.coe_zero, contDiffOn_zero] at hg ⊢ exact continuousOn_convolution_right_with_param L hk hgs hf hg | hsuc n ih => + simp only [Nat.succ_eq_add_one, Nat.cast_add, Nat.cast_one, WithTop.coe_add, + WithTop.coe_natCast, WithTop.coe_one] at hg ⊢ let f' : P → G → P × G →L[𝕜] F := fun p a => (f ⋆[L.precompR (P × G), μ] fun x : G => fderiv 𝕜 (uncurry g) (p, x)) a have A : ∀ q₀ : P × G, q₀.1 ∈ s → HasFDerivAt (fun q : P × G => (f ⋆[L, μ] g q.1) q.2) (f' q₀.1 q₀.2) q₀ := hasFDerivAt_convolution_right_with_param L hs hk hgs hf hg.one_of_succ - rw [show ((n + 1 : ℕ) : ℕ∞) = n + 1 from rfl, - contDiffOn_succ_iff_fderiv_of_isOpen (hs.prod (@isOpen_univ G _))] at hg ⊢ - constructor + rw [contDiffOn_succ_iff_fderiv_of_isOpen (hs.prod (@isOpen_univ G _))] at hg ⊢ + refine ⟨?_, ?_⟩ · rintro ⟨p, x⟩ ⟨hp, -⟩ exact (A (p, x) hp).differentiableAt.differentiableWithinAt · suffices H : ContDiffOn 𝕜 n (↿f') (s ×ˢ univ) by diff --git a/Mathlib/Analysis/Distribution/AEEqOfIntegralContDiff.lean b/Mathlib/Analysis/Distribution/AEEqOfIntegralContDiff.lean index b6d2c67574d34..8aa816c388f41 100644 --- a/Mathlib/Analysis/Distribution/AEEqOfIntegralContDiff.lean +++ b/Mathlib/Analysis/Distribution/AEEqOfIntegralContDiff.lean @@ -24,7 +24,7 @@ as `ae_eq_zero_of_integral_smooth_smul_eq_zero` and `ae_eq_of_integral_smooth_sm open MeasureTheory Filter Metric Function Set TopologicalSpace -open scoped Topology Manifold +open scoped Topology Manifold ContDiff variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [FiniteDimensional ℝ E] {F : Type*} [NormedAddCommGroup F] [NormedSpace ℝ F] [CompleteSpace F] @@ -182,7 +182,7 @@ variable [MeasurableSpace E] [BorelSpace E] {f f' : E → F} {μ : Measure E} /-- If a locally integrable function `f` on a finite-dimensional real vector space has zero integral when multiplied by any smooth compactly supported function, then `f` vanishes almost everywhere. -/ theorem ae_eq_zero_of_integral_contDiff_smul_eq_zero (hf : LocallyIntegrable f μ) - (h : ∀ (g : E → ℝ), ContDiff ℝ ⊤ g → HasCompactSupport g → ∫ x, g x • f x ∂μ = 0) : + (h : ∀ (g : E → ℝ), ContDiff ℝ ∞ g → HasCompactSupport g → ∫ x, g x • f x ∂μ = 0) : ∀ᵐ x ∂μ, f x = 0 := ae_eq_zero_of_integral_smooth_smul_eq_zero 𝓘(ℝ, E) hf (fun g g_diff g_supp ↦ h g g_diff.contDiff g_supp) @@ -192,7 +192,7 @@ integral when multiplied by any smooth compactly supported function, then they c everywhere. -/ theorem ae_eq_of_integral_contDiff_smul_eq (hf : LocallyIntegrable f μ) (hf' : LocallyIntegrable f' μ) (h : ∀ (g : E → ℝ), - ContDiff ℝ ⊤ g → HasCompactSupport g → ∫ x, g x • f x ∂μ = ∫ x, g x • f' x ∂μ) : + ContDiff ℝ ∞ g → HasCompactSupport g → ∫ x, g x • f x ∂μ = ∫ x, g x • f' x ∂μ) : ∀ᵐ x ∂μ, f x = f' x := ae_eq_of_integral_smooth_smul_eq 𝓘(ℝ, E) hf hf' (fun g g_diff g_supp ↦ h g g_diff.contDiff g_supp) @@ -202,7 +202,7 @@ theorem ae_eq_of_integral_contDiff_smul_eq in an open set `U`, then `f` vanishes almost everywhere in `U`. -/ theorem IsOpen.ae_eq_zero_of_integral_contDiff_smul_eq_zero {U : Set E} (hU : IsOpen U) (hf : LocallyIntegrableOn f U μ) - (h : ∀ (g : E → ℝ), ContDiff ℝ ⊤ g → HasCompactSupport g → tsupport g ⊆ U → + (h : ∀ (g : E → ℝ), ContDiff ℝ ∞ g → HasCompactSupport g → tsupport g ⊆ U → ∫ x, g x • f x ∂μ = 0) : ∀ᵐ x ∂μ, x ∈ U → f x = 0 := hU.ae_eq_zero_of_integral_smooth_smul_eq_zero 𝓘(ℝ, E) hf diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index 7aab3fe966d3b..b1f19e08a826b 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -61,7 +61,7 @@ Schwartz space, tempered distributions noncomputable section -open scoped Nat NNReal +open scoped Nat NNReal ContDiff variable {𝕜 𝕜' D E F G V : Type*} variable [NormedAddCommGroup E] [NormedSpace ℝ E] @@ -72,7 +72,7 @@ variable (E F) any power of `‖x‖`. -/ structure SchwartzMap where toFun : E → F - smooth' : ContDiff ℝ ⊤ toFun + smooth' : ContDiff ℝ ∞ toFun decay' : ∀ k n : ℕ, ∃ C : ℝ, ∀ x, ‖x‖ ^ k * ‖iteratedFDeriv ℝ n toFun x‖ ≤ C /-- A function is a Schwartz function if it is smooth and all derivatives decay faster than @@ -98,7 +98,7 @@ theorem decay (f : 𝓢(E, F)) (k n : ℕ) : /-- Every Schwartz function is smooth. -/ theorem smooth (f : 𝓢(E, F)) (n : ℕ∞) : ContDiff ℝ n f := - f.smooth'.of_le le_top + f.smooth'.of_le (mod_cast le_top) /-- Every Schwartz function is continuous. -/ @[continuity] @@ -521,7 +521,7 @@ section TemperateGrowth /-- A function is called of temperate growth if it is smooth and all iterated derivatives are polynomially bounded. -/ def _root_.Function.HasTemperateGrowth (f : E → F) : Prop := - ContDiff ℝ ⊤ f ∧ ∀ n : ℕ, ∃ (k : ℕ) (C : ℝ), ∀ x, ‖iteratedFDeriv ℝ n f x‖ ≤ C * (1 + ‖x‖) ^ k + ContDiff ℝ ∞ f ∧ ∀ n : ℕ, ∃ (k : ℕ) (C : ℝ), ∀ x, ‖iteratedFDeriv ℝ n f x‖ ≤ C * (1 + ‖x‖) ^ k theorem _root_.Function.HasTemperateGrowth.norm_iteratedFDeriv_le_uniform_aux {f : E → F} (hf_temperate : f.HasTemperateGrowth) (n : ℕ) : @@ -680,7 +680,7 @@ variable {σ : 𝕜 →+* 𝕜'} Note: This is a helper definition for `mkCLM`. -/ def mkLM (A : (D → E) → F → G) (hadd : ∀ (f g : 𝓢(D, E)) (x), A (f + g) x = A f x + A g x) (hsmul : ∀ (a : 𝕜) (f : 𝓢(D, E)) (x), A (a • f) x = σ a • A f x) - (hsmooth : ∀ f : 𝓢(D, E), ContDiff ℝ ⊤ (A f)) + (hsmooth : ∀ f : 𝓢(D, E), ContDiff ℝ ∞ (A f)) (hbound : ∀ n : ℕ × ℕ, ∃ (s : Finset (ℕ × ℕ)) (C : ℝ), 0 ≤ C ∧ ∀ (f : 𝓢(D, E)) (x : F), ‖x‖ ^ n.fst * ‖iteratedFDeriv ℝ n.snd (A f) x‖ ≤ C * s.sup (schwartzSeminormFamily 𝕜 D E) f) : 𝓢(D, E) →ₛₗ[σ] 𝓢(F, G) where @@ -700,7 +700,7 @@ For an example of using this definition, see `fderivCLM`. -/ def mkCLM [RingHomIsometric σ] (A : (D → E) → F → G) (hadd : ∀ (f g : 𝓢(D, E)) (x), A (f + g) x = A f x + A g x) (hsmul : ∀ (a : 𝕜) (f : 𝓢(D, E)) (x), A (a • f) x = σ a • A f x) - (hsmooth : ∀ f : 𝓢(D, E), ContDiff ℝ ⊤ (A f)) + (hsmooth : ∀ f : 𝓢(D, E), ContDiff ℝ ∞ (A f)) (hbound : ∀ n : ℕ × ℕ, ∃ (s : Finset (ℕ × ℕ)) (C : ℝ), 0 ≤ C ∧ ∀ (f : 𝓢(D, E)) (x : F), ‖x‖ ^ n.fst * ‖iteratedFDeriv ℝ n.snd (A f) x‖ ≤ C * s.sup (schwartzSeminormFamily 𝕜 D E) f) : 𝓢(D, E) →SL[σ] 𝓢(F, G) where @@ -740,19 +740,16 @@ variable [NormedField 𝕜] [NormedSpace 𝕜 F] [SMulCommClass ℝ 𝕜 F] /-- The map applying a vector to Hom-valued Schwartz function as a continuous linear map. -/ protected def evalCLM (m : E) : 𝓢(E, E →L[ℝ] F) →L[𝕜] 𝓢(E, F) := mkCLM (fun f x => f x m) (fun _ _ _ => rfl) (fun _ _ _ => rfl) - (fun f => ContDiff.clm_apply f.2 contDiff_const) - (by - rintro ⟨k, n⟩ - use {(k, n)}, ‖m‖, norm_nonneg _ - intro f x - refine - le_trans - (mul_le_mul_of_nonneg_left (norm_iteratedFDeriv_clm_apply_const f.2 le_top) - (by positivity)) - ?_ - move_mul [‖m‖] - gcongr ?_ * ‖m‖ - simp only [Finset.sup_singleton, schwartzSeminormFamily_apply, le_seminorm]) + (fun f => ContDiff.clm_apply f.2 contDiff_const) <| by + rintro ⟨k, n⟩ + use {(k, n)}, ‖m‖, norm_nonneg _ + intro f x + refine le_trans + (mul_le_mul_of_nonneg_left (norm_iteratedFDeriv_clm_apply_const f.2 (mod_cast le_top)) + (by positivity)) ?_ + move_mul [‖m‖] + gcongr ?_ * ‖m‖ + simp only [Finset.sup_singleton, schwartzSeminormFamily_apply, le_seminorm] end EvalCLM @@ -782,7 +779,8 @@ def bilinLeftCLM (B : E →L[ℝ] F →L[ℝ] G) {g : D → F} (hg : g.HasTemper intro f x have hxk : 0 ≤ ‖x‖ ^ k := by positivity have hnorm_mul := - ContinuousLinearMap.norm_iteratedFDeriv_le_of_bilinear B f.smooth' hg.1 x (n := n) le_top + ContinuousLinearMap.norm_iteratedFDeriv_le_of_bilinear B f.smooth' hg.1 x (n := n) + (mod_cast le_top) refine le_trans (mul_le_mul_of_nonneg_left hnorm_mul hxk) ?_ move_mul [← ‖B‖] simp_rw [mul_assoc ‖B‖] @@ -864,7 +862,7 @@ def compCLM {g : D → E} (hg : g.HasTemperateGrowth) · exact le_trans (by simp [hC]) (le_self_pow₀ (by simp [hC]) hN₁') · refine le_self_pow₀ (one_le_pow₀ ?_) hN₁' simp only [le_add_iff_nonneg_right, norm_nonneg] - have := norm_iteratedFDeriv_comp_le f.smooth' hg.1 le_top x hbound hgrowth' + have := norm_iteratedFDeriv_comp_le f.smooth' hg.1 (mod_cast le_top) x hbound hgrowth' have hxk : ‖x‖ ^ k ≤ (1 + ‖x‖) ^ k := pow_le_pow_left₀ (norm_nonneg _) (by simp only [zero_le_one, le_add_iff_nonneg_left]) _ refine le_trans (mul_le_mul hxk this (by positivity) (by positivity)) ?_ @@ -1018,7 +1016,8 @@ theorem iteratedPDeriv_eq_iteratedFDeriv {n : ℕ} {m : Fin n → E} {f : 𝓢(E simp only [iteratedPDeriv_succ_left, iteratedFDeriv_succ_apply_left] rw [← fderiv_continuousMultilinear_apply_const_apply] · simp [← ih] - · exact f.smooth'.differentiable_iteratedFDeriv (WithTop.coe_lt_top n) _ + · exact f.smooth'.differentiable_iteratedFDeriv (mod_cast ENat.coe_lt_top n) x + end Derivatives @@ -1053,7 +1052,7 @@ lemma integrable_pow_mul_iteratedFDeriv (f : 𝓢(D, V)) (k n : ℕ) : Integrable (fun x ↦ ‖x‖ ^ k * ‖iteratedFDeriv ℝ n f x‖) μ := integrable_of_le_of_pow_mul_le (norm_iteratedFDeriv_le_seminorm ℝ _ _) (le_seminorm ℝ _ _ _) - ((f.smooth ⊤).continuous_iteratedFDeriv le_top).aestronglyMeasurable + ((f.smooth ⊤).continuous_iteratedFDeriv (mod_cast le_top)).aestronglyMeasurable variable (μ) in lemma integrable_pow_mul (f : 𝓢(D, V)) diff --git a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean index 5b882dad07252..0151b54688137 100644 --- a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean +++ b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean @@ -81,7 +81,7 @@ noncomputable section open Real Complex MeasureTheory Filter TopologicalSpace -open scoped FourierTransform Topology +open scoped FourierTransform Topology ContDiff -- without this local instance, Lean tries first the instance -- `secondCountableTopologyEither_of_right` (whose priority is 100) and takes a very long time to @@ -310,7 +310,8 @@ lemma _root_.Continuous.fourierPowSMulRight {f : V → E} (hf : Continuous f) (n apply (smulRightL ℝ (fun (_ : Fin n) ↦ W) E).continuous₂.comp₂ _ hf exact Continuous.comp (map_continuous _) (continuous_pi (fun _ ↦ L.continuous)) -lemma _root_.ContDiff.fourierPowSMulRight {f : V → E} {k : ℕ∞} (hf : ContDiff ℝ k f) (n : ℕ) : +lemma _root_.ContDiff.fourierPowSMulRight + {f : V → E} {k : WithTop ℕ∞} (hf : ContDiff ℝ k f) (n : ℕ) : ContDiff ℝ k (fun v ↦ fourierPowSMulRight L f v n) := by simp_rw [fourierPowSMulRight_eq_comp] apply ContDiff.const_smul @@ -335,7 +336,7 @@ set_option maxSynthPendingDepth 2 in /-- The iterated derivative of a function multiplied by `(L v ⬝) ^ n` can be controlled in terms of the iterated derivatives of the initial function. -/ lemma norm_iteratedFDeriv_fourierPowSMulRight - {f : V → E} {K : ℕ∞} {C : ℝ} (hf : ContDiff ℝ K f) {n : ℕ} {k : ℕ} (hk : k ≤ K) + {f : V → E} {K : WithTop ℕ∞} {C : ℝ} (hf : ContDiff ℝ K f) {n : ℕ} {k : ℕ} (hk : k ≤ K) {v : V} (hv : ∀ i ≤ k, ∀ j ≤ n, ‖v‖ ^ j * ‖iteratedFDeriv ℝ i f v‖ ≤ C) : ‖iteratedFDeriv ℝ k (fun v ↦ fourierPowSMulRight L f v n) v‖ ≤ (2 * π) ^ n * (2 * n + 2) ^ k * ‖L‖ ^ n * C := by @@ -552,7 +553,8 @@ theorem fourierIntegral_iteratedFDeriv [FiniteDimensional ℝ V] ← fourierIntegral_continuousLinearMap_apply' J] exact H have h'n : n < N := (Nat.cast_lt.mpr n.lt_succ_self).trans_le hn - rw [fourierIntegral_fderiv _ (h'f n h'n.le) (hf.differentiable_iteratedFDeriv h'n) J] + rw [fourierIntegral_fderiv _ (h'f n h'n.le) + (hf.differentiable_iteratedFDeriv (mod_cast h'n)) J] simp only [ih h'n.le, fourierSMulRight_apply, ContinuousLinearMap.neg_apply, ContinuousLinearMap.flip_apply, neg_smul, smul_neg, neg_neg, smul_apply, fourierPowSMulRight_apply, ← coe_smul (E := E), smul_smul] @@ -583,9 +585,9 @@ theorem fourierPowSMulRight_iteratedFDeriv_fourierIntegral [FiniteDimensional simp only [Finset.mem_product, Finset.mem_range_succ_iff] at hp exact h'f _ _ ((Nat.cast_le.2 hp.1).trans hk) ((Nat.cast_le.2 hp.2).trans hm) apply (I.const_mul ((2 * π) ^ k * (2 * k + 2) ^ m * ‖L‖ ^ k)).mono' - ((hf.fourierPowSMulRight L k).continuous_iteratedFDeriv hm).aestronglyMeasurable + ((hf.fourierPowSMulRight L k).continuous_iteratedFDeriv (mod_cast hm)).aestronglyMeasurable filter_upwards with v - refine norm_iteratedFDeriv_fourierPowSMulRight _ hf hm (fun i hi j hj ↦ ?_) + refine norm_iteratedFDeriv_fourierPowSMulRight _ hf (mod_cast hm) (fun i hi j hj ↦ ?_) apply Finset.single_le_sum (f := fun p ↦ ‖v‖ ^ p.1 * ‖iteratedFDeriv ℝ p.2 f v‖) (a := (j, i)) · intro i _hi positivity @@ -616,7 +618,7 @@ theorem norm_fourierPowSMulRight_iteratedFDeriv_fourierIntegral_le [FiniteDimens · filter_upwards with v using norm_nonneg _ · exact (integrable_finset_sum _ I).const_mul _ · filter_upwards with v - apply norm_iteratedFDeriv_fourierPowSMulRight _ hf hn _ + apply norm_iteratedFDeriv_fourierPowSMulRight _ hf (mod_cast hn) _ intro i hi j hj apply Finset.single_le_sum (f := fun p ↦ ‖v‖ ^ p.1 * ‖iteratedFDeriv ℝ p.2 f v‖) (a := (j, i)) · intro i _hi diff --git a/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean b/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean index 3a485d339830b..846339ecb2f00 100644 --- a/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean @@ -90,7 +90,7 @@ 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 := +theorem contDiff_exp {n : WithTop ℕ∞} : ContDiff 𝕜 n exp := analyticOnNhd_cexp.restrictScalars.contDiff theorem hasStrictDerivAt_exp (x : ℂ) : HasStrictDerivAt exp (exp x) x := @@ -230,7 +230,7 @@ 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 : WithTop ℕ∞} : ContDiff ℝ n exp := Complex.contDiff_exp.real_of_complex @[simp] diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean index e9c7ca6362653..b886b2c966010 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean @@ -24,7 +24,7 @@ logarithm, derivative open Filter Finset Set -open scoped Topology +open scoped Topology ContDiff namespace Real @@ -66,12 +66,13 @@ theorem deriv_log (x : ℝ) : deriv log x = x⁻¹ := theorem deriv_log' : deriv log = Inv.inv := funext deriv_log -theorem contDiffOn_log {n : ℕ∞} : ContDiffOn ℝ n log {0}ᶜ := by - suffices ContDiffOn ℝ ⊤ log {0}ᶜ from this.of_le le_top +theorem contDiffOn_log {n : WithTop ℕ∞} : ContDiffOn ℝ n log {0}ᶜ := by + suffices ContDiffOn ℝ ω log {0}ᶜ from this.of_le le_top + rw [← contDiffOn_infty_iff_contDiffOn_omega] refine (contDiffOn_top_iff_deriv_of_isOpen isOpen_compl_singleton).2 ?_ simp [differentiableOn_log, contDiffOn_inv] -theorem contDiffAt_log {n : ℕ∞} : ContDiffAt ℝ n log x ↔ x ≠ 0 := +theorem contDiffAt_log {n : WithTop ℕ∞} : ContDiffAt ℝ n log x ↔ x ≠ 0 := ⟨fun h => continuousAt_log_iff.1 h.continuousAt, fun hx => (contDiffOn_log x hx).contDiffAt <| IsOpen.mem_nhds isOpen_compl_singleton hx⟩ diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean index 49c7cb8a86df9..0c619bd07f671 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean @@ -282,7 +282,7 @@ theorem hasStrictFDerivAt_rpow_of_neg (p : ℝ × ℝ) (hp : p.1 < 0) : rw [div_eq_mul_inv, add_comm]; congr 2 <;> ring /-- The function `fun (x, y) => x ^ y` is infinitely smooth at `(x, y)` unless `x = 0`. -/ -theorem contDiffAt_rpow_of_ne (p : ℝ × ℝ) (hp : p.1 ≠ 0) {n : ℕ∞} : +theorem contDiffAt_rpow_of_ne (p : ℝ × ℝ) (hp : p.1 ≠ 0) {n : WithTop ℕ∞} : ContDiffAt ℝ n (fun p : ℝ × ℝ => p.1 ^ p.2) p := by cases' hp.lt_or_lt with hneg hpos exacts @@ -358,7 +358,7 @@ theorem deriv_rpow_const' {p : ℝ} (h : 1 ≤ p) : (deriv fun x : ℝ => x ^ p) = fun x => p * x ^ (p - 1) := funext fun _ => deriv_rpow_const (Or.inr h) -theorem contDiffAt_rpow_const_of_ne {x p : ℝ} {n : ℕ∞} (h : x ≠ 0) : +theorem contDiffAt_rpow_const_of_ne {x p : ℝ} {n : WithTop ℕ∞} (h : x ≠ 0) : ContDiffAt ℝ n (fun x => x ^ p) x := (contDiffAt_rpow_of_ne (x, p) h).comp x (contDiffAt_id.prod contDiffAt_const) @@ -368,7 +368,8 @@ theorem contDiff_rpow_const_of_le {p : ℝ} {n : ℕ} (h : ↑n ≤ p) : · exact contDiff_zero.2 (continuous_id.rpow_const fun x => Or.inr <| by simpa using h) · have h1 : 1 ≤ p := le_trans (by simp) h rw [Nat.cast_succ, ← le_sub_iff_add_le] at h - rw [show ((n + 1 : ℕ) : ℕ∞) = n + 1 from rfl, contDiff_succ_iff_deriv, deriv_rpow_const' h1] + rw [show ((n + 1 : ℕ) : WithTop ℕ∞) = n + 1 from rfl, contDiff_succ_iff_deriv, + deriv_rpow_const' h1] exact ⟨differentiable_rpow_const h1, contDiff_const.mul (ihn h)⟩ theorem contDiffAt_rpow_const_of_le {x p : ℝ} {n : ℕ} (h : ↑n ≤ p) : diff --git a/Mathlib/Analysis/SpecialFunctions/SmoothTransition.lean b/Mathlib/Analysis/SpecialFunctions/SmoothTransition.lean index 20b9a0ce94d94..a9815ec827dfa 100644 --- a/Mathlib/Analysis/SpecialFunctions/SmoothTransition.lean +++ b/Mathlib/Analysis/SpecialFunctions/SmoothTransition.lean @@ -113,7 +113,7 @@ theorem contDiff_polynomial_eval_inv_mul {n : ℕ∞} (p : ℝ[X]) : exact (hasDerivAt_polynomial_eval_inv_mul p _).deriv /-- The function `expNegInvGlue` is smooth. -/ -protected theorem contDiff {n} : ContDiff ℝ n expNegInvGlue := by +protected theorem contDiff {n : ℕ∞} : ContDiff ℝ n expNegInvGlue := by simpa using contDiff_polynomial_eval_inv_mul 1 end expNegInvGlue @@ -174,12 +174,12 @@ theorem lt_one_of_lt_one (h : x < 1) : smoothTransition x < 1 := theorem pos_of_pos (h : 0 < x) : 0 < smoothTransition x := div_pos (expNegInvGlue.pos_of_pos h) (pos_denom x) -protected theorem contDiff {n} : ContDiff ℝ n smoothTransition := +protected theorem contDiff {n : ℕ∞} : ContDiff ℝ n smoothTransition := expNegInvGlue.contDiff.div (expNegInvGlue.contDiff.add <| expNegInvGlue.contDiff.comp <| contDiff_const.sub contDiff_id) fun x => (pos_denom x).ne' -protected theorem contDiffAt {x n} : ContDiffAt ℝ n smoothTransition x := +protected theorem contDiffAt {x : ℝ} {n : ℕ∞} : ContDiffAt ℝ n smoothTransition x := smoothTransition.contDiff.contDiffAt protected theorem continuous : Continuous smoothTransition := diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/ComplexDeriv.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/ComplexDeriv.lean index e414b1421381a..a06b4a537e039 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/ComplexDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/ComplexDeriv.lean @@ -62,7 +62,7 @@ theorem deriv_tan (x : ℂ) : deriv tan x = 1 / cos x ^ 2 := else (hasDerivAt_tan h).deriv @[simp] -theorem contDiffAt_tan {x : ℂ} {n : ℕ∞} : ContDiffAt ℂ n tan x ↔ cos x ≠ 0 := +theorem contDiffAt_tan {x : ℂ} {n : WithTop ℕ∞} : ContDiffAt ℂ n tan x ↔ cos x ≠ 0 := ⟨fun h => continuousAt_tan.1 h.continuousAt, contDiff_sin.contDiffAt.div contDiff_cos.contDiffAt⟩ end Complex diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/InverseDeriv.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/InverseDeriv.lean index 6ade7f02ca336..d9efa5b942979 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/InverseDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/InverseDeriv.lean @@ -14,7 +14,7 @@ Derivatives of `arcsin` and `arccos`. noncomputable section -open scoped Topology Filter Real +open scoped Topology Filter Real ContDiff open Set namespace Real @@ -22,7 +22,7 @@ namespace Real section Arcsin theorem deriv_arcsin_aux {x : ℝ} (h₁ : x ≠ -1) (h₂ : x ≠ 1) : - HasStrictDerivAt arcsin (1 / √(1 - x ^ 2)) x ∧ ContDiffAt ℝ ⊤ arcsin x := by + HasStrictDerivAt arcsin (1 / √(1 - x ^ 2)) x ∧ ContDiffAt ℝ ω arcsin x := by cases' h₁.lt_or_lt with h₁ h₁ · have : 1 - x ^ 2 < 0 := by nlinarith [h₁] rw [sqrt_eq_zero'.2 this.le, div_zero] @@ -50,7 +50,8 @@ theorem hasDerivAt_arcsin {x : ℝ} (h₁ : x ≠ -1) (h₂ : x ≠ 1) : HasDerivAt arcsin (1 / √(1 - x ^ 2)) x := (hasStrictDerivAt_arcsin h₁ h₂).hasDerivAt -theorem contDiffAt_arcsin {x : ℝ} (h₁ : x ≠ -1) (h₂ : x ≠ 1) {n : ℕ∞} : ContDiffAt ℝ n arcsin x := +theorem contDiffAt_arcsin {x : ℝ} (h₁ : x ≠ -1) (h₂ : x ≠ 1) {n : WithTop ℕ∞} : + ContDiffAt ℝ n arcsin x := (deriv_arcsin_aux h₁ h₂).2.of_le le_top theorem hasDerivWithinAt_arcsin_Ici {x : ℝ} (h : x ≠ -1) : @@ -102,12 +103,13 @@ theorem differentiableOn_arcsin : DifferentiableOn ℝ arcsin {-1, 1}ᶜ := fun (differentiableAt_arcsin.2 ⟨fun h => hx (Or.inl h), fun h => hx (Or.inr h)⟩).differentiableWithinAt -theorem contDiffOn_arcsin {n : ℕ∞} : ContDiffOn ℝ n arcsin {-1, 1}ᶜ := fun _x hx => +theorem contDiffOn_arcsin {n : WithTop ℕ∞} : ContDiffOn ℝ n arcsin {-1, 1}ᶜ := fun _x hx => (contDiffAt_arcsin (mt Or.inl hx) (mt Or.inr hx)).contDiffWithinAt -theorem contDiffAt_arcsin_iff {x : ℝ} {n : ℕ∞} : ContDiffAt ℝ n arcsin x ↔ n = 0 ∨ x ≠ -1 ∧ x ≠ 1 := +theorem contDiffAt_arcsin_iff {x : ℝ} {n : WithTop ℕ∞} : + ContDiffAt ℝ n arcsin x ↔ n = 0 ∨ x ≠ -1 ∧ x ≠ 1 := ⟨fun h => or_iff_not_imp_left.2 fun hn => differentiableAt_arcsin.1 <| h.differentiableAt <| - ENat.one_le_iff_ne_zero.2 hn, + ENat.one_le_iff_ne_zero_withTop.mpr hn, fun h => h.elim (fun hn => hn.symm ▸ (contDiff_zero.2 continuous_arcsin).contDiffAt) fun hx => contDiffAt_arcsin hx.1 hx.2⟩ @@ -123,7 +125,8 @@ theorem hasDerivAt_arccos {x : ℝ} (h₁ : x ≠ -1) (h₂ : x ≠ 1) : HasDerivAt arccos (-(1 / √(1 - x ^ 2))) x := (hasDerivAt_arcsin h₁ h₂).const_sub (π / 2) -theorem contDiffAt_arccos {x : ℝ} (h₁ : x ≠ -1) (h₂ : x ≠ 1) {n : ℕ∞} : ContDiffAt ℝ n arccos x := +theorem contDiffAt_arccos {x : ℝ} (h₁ : x ≠ -1) (h₂ : x ≠ 1) {n : WithTop ℕ∞} : + ContDiffAt ℝ n arccos x := contDiffAt_const.sub (contDiffAt_arcsin h₁ h₂) theorem hasDerivWithinAt_arccos_Ici {x : ℝ} (h : x ≠ -1) : @@ -152,10 +155,10 @@ theorem deriv_arccos : deriv arccos = fun x => -(1 / √(1 - x ^ 2)) := theorem differentiableOn_arccos : DifferentiableOn ℝ arccos {-1, 1}ᶜ := differentiableOn_arcsin.const_sub _ -theorem contDiffOn_arccos {n : ℕ∞} : ContDiffOn ℝ n arccos {-1, 1}ᶜ := +theorem contDiffOn_arccos {n : WithTop ℕ∞} : ContDiffOn ℝ n arccos {-1, 1}ᶜ := contDiffOn_const.sub contDiffOn_arcsin -theorem contDiffAt_arccos_iff {x : ℝ} {n : ℕ∞} : +theorem contDiffAt_arccos_iff {x : ℝ} {n : WithTop ℕ∞} : ContDiffAt ℝ n arccos x ↔ n = 0 ∨ x ≠ -1 ∧ x ≠ 1 := by refine Iff.trans ⟨fun h => ?_, fun h => ?_⟩ contDiffAt_arcsin_iff <;> simpa [arccos] using (contDiffAt_const (c := π / 2)).sub h diff --git a/Mathlib/Geometry/Manifold/Algebra/LeftInvariantDerivation.lean b/Mathlib/Geometry/Manifold/Algebra/LeftInvariantDerivation.lean index 998cc71804f14..9e2358ae486b1 100644 --- a/Mathlib/Geometry/Manifold/Algebra/LeftInvariantDerivation.lean +++ b/Mathlib/Geometry/Manifold/Algebra/LeftInvariantDerivation.lean @@ -23,6 +23,9 @@ implementing one of the possible definitions of the Lie algebra attached to a Li noncomputable section open scoped LieGroup Manifold Derivation +/- Next line is necessary while the manifold smoothness class is not extended to `ω`. +Later, replace with `open scoped ContDiff`. -/ +local notation "∞" => (⊤ : ℕ∞) variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) (G : Type*) diff --git a/Mathlib/Geometry/Manifold/Algebra/Monoid.lean b/Mathlib/Geometry/Manifold/Algebra/Monoid.lean index 3808ac40a51c1..7e6a90b613a22 100644 --- a/Mathlib/Geometry/Manifold/Algebra/Monoid.lean +++ b/Mathlib/Geometry/Manifold/Algebra/Monoid.lean @@ -18,6 +18,9 @@ semigroups. open scoped Manifold +/- Next line is necessary while the manifold smoothness class is not extended to `ω`. +Later, replace with `open scoped ContDiff`. -/ +local notation "∞" => (⊤ : ℕ∞) library_note "Design choices about smooth algebraic structures"/-- 1. All smooth algebraic structures on `G` are `Prop`-valued classes that extend diff --git a/Mathlib/Geometry/Manifold/Algebra/SmoothFunctions.lean b/Mathlib/Geometry/Manifold/Algebra/SmoothFunctions.lean index e8adb09c03d94..db675c91210d4 100644 --- a/Mathlib/Geometry/Manifold/Algebra/SmoothFunctions.lean +++ b/Mathlib/Geometry/Manifold/Algebra/SmoothFunctions.lean @@ -15,6 +15,9 @@ In this file, we define instances of algebraic structures over smooth functions. noncomputable section open scoped Manifold +/- Next line is necessary while the manifold smoothness class is not extended to `ω`. +Later, replace with `open scoped ContDiff`. -/ +local notation "∞" => (⊤ : ℕ∞) open TopologicalSpace diff --git a/Mathlib/Geometry/Manifold/Algebra/Structures.lean b/Mathlib/Geometry/Manifold/Algebra/Structures.lean index da593ec373f61..a20f6b529acf8 100644 --- a/Mathlib/Geometry/Manifold/Algebra/Structures.lean +++ b/Mathlib/Geometry/Manifold/Algebra/Structures.lean @@ -12,7 +12,7 @@ In this file we define smooth structures that build on Lie groups. We prefer usi instead of Lie mainly because Lie ring has currently another use in mathematics. -/ -open scoped Manifold +open scoped Manifold ContDiff section SmoothRing @@ -35,7 +35,7 @@ instance (priority := 100) SmoothRing.toSmoothMul (I : ModelWithCorners 𝕜 E H -- see Note [lower instance priority] instance (priority := 100) SmoothRing.toLieAddGroup (I : ModelWithCorners 𝕜 E H) (R : Type*) [Ring R] [TopologicalSpace R] [ChartedSpace H R] [SmoothRing I R] : LieAddGroup I R where - compatible := StructureGroupoid.compatible (contDiffGroupoid ⊤ I) + compatible := StructureGroupoid.compatible (contDiffGroupoid ∞ I) smooth_add := contMDiff_add I smooth_neg := by simpa only [neg_one_mul] using contMDiff_mul_left (G := R) (a := -1) diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean b/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean index 2036cda540d21..31ff5e240b2d7 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean @@ -13,7 +13,7 @@ and that local structomorphisms are smooth with smooth inverses. -/ open Set ChartedSpace SmoothManifoldWithCorners -open scoped Manifold +open scoped Manifold ContDiff variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] -- declare a smooth manifold `M` over the pair `(E, H)`. @@ -46,18 +46,14 @@ theorem contMDiffOn_model_symm : ContMDiffOn 𝓘(𝕜, E) I n I.symm (range I) /-- An atlas member is `C^n` for any `n`. -/ theorem contMDiffOn_of_mem_maximalAtlas (h : e ∈ maximalAtlas I M) : ContMDiffOn I I n e e.source := - ContMDiffOn.of_le - ((contDiffWithinAt_localInvariantProp ∞).liftPropOn_of_mem_maximalAtlas - contDiffWithinAtProp_id h) - le_top + ContMDiffOn.of_le ((contDiffWithinAt_localInvariantProp ⊤).liftPropOn_of_mem_maximalAtlas + contDiffWithinAtProp_id h) le_top /-- The inverse of an atlas member is `C^n` for any `n`. -/ theorem contMDiffOn_symm_of_mem_maximalAtlas (h : e ∈ maximalAtlas I M) : ContMDiffOn I I n e.symm e.target := - ContMDiffOn.of_le - ((contDiffWithinAt_localInvariantProp ∞).liftPropOn_symm_of_mem_maximalAtlas - contDiffWithinAtProp_id h) - le_top + ContMDiffOn.of_le ((contDiffWithinAt_localInvariantProp ⊤).liftPropOn_symm_of_mem_maximalAtlas + contDiffWithinAtProp_id h) le_top theorem contMDiffAt_of_mem_maximalAtlas (h : e ∈ maximalAtlas I M) (hx : x ∈ e.source) : ContMDiffAt I I n e x := @@ -112,7 +108,7 @@ theorem contMDiffWithinAt_extChartAt_symm_range /-- An element of `contDiffGroupoid ⊤ I` is `C^n` for any `n`. -/ theorem contMDiffOn_of_mem_contDiffGroupoid {e' : PartialHomeomorph H H} - (h : e' ∈ contDiffGroupoid ⊤ I) : ContMDiffOn I I n e' e'.source := + (h : e' ∈ contDiffGroupoid ∞ I) : ContMDiffOn I I n e' e'.source := (contDiffWithinAt_localInvariantProp n).liftPropOn_of_mem_groupoid contDiffWithinAtProp_id h end Atlas @@ -124,7 +120,7 @@ section IsLocalStructomorph variable [ChartedSpace H M'] [IsM' : SmoothManifoldWithCorners I M'] theorem isLocalStructomorphOn_contDiffGroupoid_iff_aux {f : PartialHomeomorph M M'} - (hf : LiftPropOn (contDiffGroupoid ⊤ I).IsLocalStructomorphWithinAt f f.source) : + (hf : LiftPropOn (contDiffGroupoid ∞ I).IsLocalStructomorphWithinAt f f.source) : ContMDiffOn I I ⊤ f f.source := by -- It suffices to show smoothness near each `x` apply contMDiffOn_of_locally_contMDiffOn @@ -170,7 +166,7 @@ theorem isLocalStructomorphOn_contDiffGroupoid_iff_aux {f : PartialHomeomorph M is a local structomorphism for `I`, if and only if it is manifold-smooth on the domain of definition in both directions. -/ theorem isLocalStructomorphOn_contDiffGroupoid_iff (f : PartialHomeomorph M M') : - LiftPropOn (contDiffGroupoid ⊤ I).IsLocalStructomorphWithinAt f f.source ↔ + LiftPropOn (contDiffGroupoid ∞ I).IsLocalStructomorphWithinAt f f.source ↔ ContMDiffOn I I ⊤ f f.source ∧ ContMDiffOn I I ⊤ f.symm f.target := by constructor · intro h @@ -186,7 +182,7 @@ theorem isLocalStructomorphOn_contDiffGroupoid_iff (f : PartialHomeomorph M M') refine ⟨(f.symm.continuousAt hX).continuousWithinAt, fun h2x => ?_⟩ obtain ⟨e, he, h2e, hef, hex⟩ : ∃ e : PartialHomeomorph H H, - e ∈ contDiffGroupoid ⊤ I ∧ + e ∈ contDiffGroupoid ∞ I ∧ e.source ⊆ (c.symm ≫ₕ f ≫ₕ c').source ∧ EqOn (c' ∘ f ∘ c.symm) e e.source ∧ c x ∈ e.source := by have h1 : c' = chartAt H (f x) := by simp only [f.right_inv hX] diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean b/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean index f5506bc51360d..bbe1eedcdb768 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean @@ -156,7 +156,7 @@ section id theorem contMDiff_id : ContMDiff I I n (id : M → M) := ContMDiff.of_le - ((contDiffWithinAt_localInvariantProp ∞).liftProp_id contDiffWithinAtProp_id) le_top + ((contDiffWithinAt_localInvariantProp ⊤).liftProp_id contDiffWithinAtProp_id) le_top @[deprecated (since := "2024-11-20")] alias smooth_id := contMDiff_id diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean index aa5e8d957cfd7..49996ee7ba8cb 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean @@ -46,7 +46,7 @@ in terms of extended charts in `contMDiffOn_iff` and `contMDiff_iff`. open Set Function Filter ChartedSpace SmoothManifoldWithCorners -open scoped Topology Manifold +open scoped Topology Manifold ContDiff /-! ### Definition of smooth functions between manifolds -/ @@ -110,7 +110,8 @@ theorem contDiffWithinAt_localInvariantProp (n : ℕ∞) : rw [this] at h have : I (e x) ∈ I.symm ⁻¹' e.target ∩ range I := by simp only [hx, mfld_simps] have := (mem_groupoid_of_pregroupoid.2 he).2.contDiffWithinAt this - convert (h.comp_inter _ (this.of_le le_top)).mono_of_mem_nhdsWithin _ using 1 + convert (h.comp_inter _ (this.of_le (mod_cast le_top))).mono_of_mem_nhdsWithin _ + using 1 · ext y; simp only [mfld_simps] refine mem_nhdsWithin.mpr ⟨I.symm ⁻¹' e.target, e.open_target.preimage I.continuous_symm, by @@ -127,7 +128,7 @@ theorem contDiffWithinAt_localInvariantProp (n : ℕ∞) : have A : (I' ∘ f ∘ I.symm) (I x) ∈ I'.symm ⁻¹' e'.source ∩ range I' := by simp only [hx, mfld_simps] have := (mem_groupoid_of_pregroupoid.2 he').1.contDiffWithinAt A - convert (this.of_le le_top).comp _ h _ + convert (this.of_le (mod_cast le_top)).comp _ h _ · ext y; simp only [mfld_simps] · intro y hy; simp only [mfld_simps] at hy; simpa only [hy, mfld_simps] using hs hy.1 @@ -197,7 +198,7 @@ def ContMDiff (n : ℕ∞) (f : M → M') := theorem ContMDiffWithinAt.of_le (hf : ContMDiffWithinAt I I' n f s x) (le : m ≤ n) : ContMDiffWithinAt I I' m f s x := by simp only [ContMDiffWithinAt, LiftPropWithinAt] at hf ⊢ - exact ⟨hf.1, hf.2.of_le le⟩ + exact ⟨hf.1, hf.2.of_le (mod_cast le)⟩ theorem ContMDiffAt.of_le (hf : ContMDiffAt I I' n f x) (le : m ≤ n) : ContMDiffAt I I' m f x := ContMDiffWithinAt.of_le hf le diff --git a/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean b/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean index ad3c237062762..38b8fd8f39f73 100644 --- a/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean +++ b/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean @@ -113,7 +113,7 @@ protected theorem ContMDiffWithinAt.mfderivWithin {x₀ : N} {f : N → M → M' rw [contMDiffWithinAt_iff] at hf' hg' simp_rw [Function.comp_def, uncurry, extChartAt_prod, PartialEquiv.prod_coe_symm, ModelWithCorners.range_prod] at hf' ⊢ - apply ContDiffWithinAt.fderivWithin _ _ _ hmn + apply ContDiffWithinAt.fderivWithin _ _ _ (show (m : WithTop ℕ∞) + 1 ≤ n from mod_cast hmn ) · simp [hx₀, t'] · apply inter_subset_left.trans rw [preimage_subset_iff] @@ -404,6 +404,9 @@ namespace ContMDiffMap -- (However as a consequence we import `Mathlib/Geometry/Manifold/ContMDiffMap.lean` here now.) -- They could be moved to another file (perhaps a new file) if desired. open scoped Manifold +/- Next line is necessary while the manifold smoothness class is not extended to `ω`. +Later, replace with `open scoped ContDiff`. -/ +local notation "∞" => (⊤ : ℕ∞) protected theorem mdifferentiable' (f : C^n⟮I, M; I', M'⟯) (hn : 1 ≤ n) : MDifferentiable I I' f := f.contMDiff.mdifferentiable hn diff --git a/Mathlib/Geometry/Manifold/ContMDiffMap.lean b/Mathlib/Geometry/Manifold/ContMDiffMap.lean index f8bd9f7980ad8..2be46a551a705 100644 --- a/Mathlib/Geometry/Manifold/ContMDiffMap.lean +++ b/Mathlib/Geometry/Manifold/ContMDiffMap.lean @@ -39,6 +39,10 @@ scoped[Manifold] notation "C^" n "⟮" I ", " M "; " k "⟯" => ContMDiffMap I (modelWithCornersSelf k k) M k n open scoped Manifold +/- Next line is necessary while the manifold smoothness class is not extended to `ω`. +Later, replace with `open scoped ContDiff`. -/ +local notation "∞" => (⊤ : ℕ∞) + namespace ContMDiffMap diff --git a/Mathlib/Geometry/Manifold/DerivationBundle.lean b/Mathlib/Geometry/Manifold/DerivationBundle.lean index ca029181286c4..fb3dbd766186a 100644 --- a/Mathlib/Geometry/Manifold/DerivationBundle.lean +++ b/Mathlib/Geometry/Manifold/DerivationBundle.lean @@ -25,6 +25,9 @@ variable (𝕜 : Type*) [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCom [TopologicalSpace M] [ChartedSpace H M] (n : ℕ∞) open scoped Manifold +/- Next line is necessary while the manifold smoothness class is not extended to `ω`. +Later, replace with `open scoped ContDiff`. -/ +local notation "∞" => (⊤ : ℕ∞) -- the following two instances prevent poorly understood type class inference timeout problems instance smoothFunctionsAlgebra : Algebra 𝕜 C^∞⟮I, M; 𝕜⟯ := by infer_instance diff --git a/Mathlib/Geometry/Manifold/Diffeomorph.lean b/Mathlib/Geometry/Manifold/Diffeomorph.lean index 5ee61a8605fe2..6a6dd316cf875 100644 --- a/Mathlib/Geometry/Manifold/Diffeomorph.lean +++ b/Mathlib/Geometry/Manifold/Diffeomorph.lean @@ -44,7 +44,7 @@ diffeomorphism, manifold -/ -open scoped Manifold Topology +open scoped Manifold Topology ContDiff open Function Set @@ -491,7 +491,7 @@ instance smoothManifoldWithCorners_transDiffeomorph [SmoothManifoldWithCorners I SmoothManifoldWithCorners (I.transDiffeomorph e) M := by refine smoothManifoldWithCorners_of_contDiffOn (I.transDiffeomorph e) M fun e₁ e₂ h₁ h₂ => ?_ refine e.contDiff.comp_contDiffOn - (((contDiffGroupoid ⊤ I).compatible h₁ h₂).1.comp e.symm.contDiff.contDiffOn ?_) + (((contDiffGroupoid ∞ I).compatible h₁ h₂).1.comp e.symm.contDiff.contDiffOn ?_) simp only [mapsTo_iff_subset_preimage] mfld_set_tac diff --git a/Mathlib/Geometry/Manifold/Instances/Real.lean b/Mathlib/Geometry/Manifold/Instances/Real.lean index 26caf2f7bee47..ce5fb5a376760 100644 --- a/Mathlib/Geometry/Manifold/Instances/Real.lean +++ b/Mathlib/Geometry/Manifold/Instances/Real.lean @@ -42,7 +42,7 @@ noncomputable section open Set Function -open scoped Manifold +open scoped Manifold ContDiff /-- The half-space in `ℝ^n`, used to model manifolds with boundary. We only define it when `1 ≤ n`, as the definition only makes sense in this case. diff --git a/Mathlib/Geometry/Manifold/Instances/Sphere.lean b/Mathlib/Geometry/Manifold/Instances/Sphere.lean index 00100a2fdd4ed..a426774e69c58 100644 --- a/Mathlib/Geometry/Manifold/Instances/Sphere.lean +++ b/Mathlib/Geometry/Manifold/Instances/Sphere.lean @@ -69,7 +69,7 @@ noncomputable section open Metric Module Function -open scoped Manifold +open scoped Manifold ContDiff section StereographicProjection @@ -93,7 +93,7 @@ theorem stereoToFun_apply (x : E) : rfl theorem contDiffOn_stereoToFun : - ContDiffOn ℝ ⊤ (stereoToFun v) {x : E | innerSL _ v x ≠ (1 : ℝ)} := by + ContDiffOn ℝ ∞ (stereoToFun v) {x : E | innerSL _ v x ≠ (1 : ℝ)} := by refine ContDiffOn.smul ?_ (orthogonalProjection (ℝ ∙ v)ᗮ).contDiff.contDiffOn refine contDiff_const.contDiffOn.div ?_ ?_ · exact (contDiff_const.sub (innerSL ℝ v).contDiff).contDiffOn @@ -157,13 +157,13 @@ theorem hasFDerivAt_stereoInvFunAux_comp_coe (v : E) : hasFDerivAt_stereoInvFunAux v refine this.comp (0 : (ℝ ∙ v)ᗮ) (by apply ContinuousLinearMap.hasFDerivAt) -theorem contDiff_stereoInvFunAux : ContDiff ℝ ⊤ (stereoInvFunAux v) := by - have h₀ : ContDiff ℝ ⊤ fun w : E => ‖w‖ ^ 2 := contDiff_norm_sq ℝ - have h₁ : ContDiff ℝ ⊤ fun w : E => (‖w‖ ^ 2 + 4)⁻¹ := by +theorem contDiff_stereoInvFunAux : ContDiff ℝ ∞ (stereoInvFunAux v) := by + have h₀ : ContDiff ℝ ∞ fun w : E => ‖w‖ ^ 2 := contDiff_norm_sq ℝ + have h₁ : ContDiff ℝ ∞ fun w : E => (‖w‖ ^ 2 + 4)⁻¹ := by refine (h₀.add contDiff_const).inv ?_ intro x - positivity - have h₂ : ContDiff ℝ ⊤ fun w => (4 : ℝ) • w + (‖w‖ ^ 2 - 4) • v := by + nlinarith + 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 exact h₁.smul h₂ @@ -389,7 +389,7 @@ instance EuclideanSpace.instSmoothManifoldWithCornersSphere {n : ℕ} [Fact (fin -- Porting note: need to help with implicit variables again have H₂ := (contDiff_stereoInvFunAux (v := v.val)|>.comp (ℝ ∙ (v : E))ᗮ.subtypeL.contDiff).comp U.symm.contDiff - convert H₁.comp_inter (H₂.contDiffOn : ContDiffOn ℝ ⊤ _ Set.univ) using 1 + convert H₁.comp_inter (H₂.contDiffOn : ContDiffOn ℝ ∞ _ Set.univ) using 1 -- -- squeezed from `ext, simp [sphere_ext_iff, stereographic'_symm_apply, real_inner_comm]` simp only [PartialHomeomorph.trans_toPartialEquiv, PartialHomeomorph.symm_toPartialEquiv, PartialEquiv.trans_source, PartialEquiv.symm_source, stereographic'_target, @@ -407,7 +407,7 @@ instance (n : ℕ) : /-- The inclusion map (i.e., `coe`) from the sphere in `E` to `E` is smooth. -/ theorem contMDiff_coe_sphere {n : ℕ} [Fact (finrank ℝ E = n + 1)] : - ContMDiff (𝓡 n) 𝓘(ℝ, E) ∞ ((↑) : sphere (0 : E) 1 → E) := by + ContMDiff (𝓡 n) 𝓘(ℝ, E) ⊤ ((↑) : sphere (0 : E) 1 → E) := by -- Porting note: trouble with filling these implicit variables in the instance have := EuclideanSpace.instSmoothManifoldWithCornersSphere (E := E) (n := n) rw [contMDiff_iff] @@ -438,7 +438,7 @@ theorem ContMDiff.codRestrict_sphere {n : ℕ} [Fact (finrank ℝ E = n + 1)] {m (-- Again, partially removing type ascription... Weird that this helps! OrthonormalBasis.fromOrthogonalSpanSingleton n (ne_zero_of_mem_unit_sphere (-v))).repr - have h : ContDiffOn ℝ ⊤ _ Set.univ := U.contDiff.contDiffOn + have h : ContDiffOn ℝ ∞ _ Set.univ := U.contDiff.contDiffOn have H₁ := (h.comp_inter contDiffOn_stereoToFun).contMDiffOn have H₂ : ContMDiffOn _ _ _ _ Set.univ := hf.contMDiffOn convert (H₁.of_le le_top).comp' H₂ using 1 @@ -453,7 +453,7 @@ theorem ContMDiff.codRestrict_sphere {n : ℕ} [Fact (finrank ℝ E = n + 1)] {m /-- The antipodal map is smooth. -/ theorem contMDiff_neg_sphere {n : ℕ} [Fact (finrank ℝ E = n + 1)] : - ContMDiff (𝓡 n) (𝓡 n) ∞ fun x : sphere (0 : E) 1 => -x := by + ContMDiff (𝓡 n) (𝓡 n) ⊤ fun x : sphere (0 : E) 1 => -x := by -- this doesn't elaborate well in term mode apply ContMDiff.codRestrict_sphere apply contDiff_neg.contMDiff.comp _ @@ -555,7 +555,7 @@ instance : LieGroup (𝓡 1) Circle where smooth_mul := by apply ContMDiff.codRestrict_sphere let c : Circle → ℂ := (↑) - have h₂ : ContMDiff (𝓘(ℝ, ℂ).prod 𝓘(ℝ, ℂ)) 𝓘(ℝ, ℂ) ∞ fun z : ℂ × ℂ => z.fst * z.snd := by + have h₂ : ContMDiff (𝓘(ℝ, ℂ).prod 𝓘(ℝ, ℂ)) 𝓘(ℝ, ℂ) ⊤ fun z : ℂ × ℂ => z.fst * z.snd := by rw [contMDiff_iff] exact ⟨continuous_mul, fun x y => contDiff_mul.contDiffOn⟩ -- Porting note: needed to fill in first 3 arguments or could not figure out typeclasses @@ -569,7 +569,7 @@ instance : LieGroup (𝓡 1) Circle where exact Complex.conjCLE.contDiff.contMDiff.comp contMDiff_coe_sphere /-- The map `fun t ↦ exp (t * I)` from `ℝ` to the unit circle in `ℂ` is smooth. -/ -theorem contMDiff_circleExp : ContMDiff 𝓘(ℝ, ℝ) (𝓡 1) ∞ Circle.exp := +theorem contMDiff_circleExp : ContMDiff 𝓘(ℝ, ℝ) (𝓡 1) ⊤ Circle.exp := (contDiff_exp.comp (contDiff_id.smul contDiff_const)).contMDiff.codRestrict_sphere _ @[deprecated (since := "2024-07-25")] alias contMDiff_expMapCircle := contMDiff_circleExp diff --git a/Mathlib/Geometry/Manifold/Instances/UnitsOfNormedAlgebra.lean b/Mathlib/Geometry/Manifold/Instances/UnitsOfNormedAlgebra.lean index 8ee99e793fb7c..b66b8302d348a 100644 --- a/Mathlib/Geometry/Manifold/Instances/UnitsOfNormedAlgebra.lean +++ b/Mathlib/Geometry/Manifold/Instances/UnitsOfNormedAlgebra.lean @@ -27,6 +27,9 @@ example {V : Type*} [NormedAddCommGroup V] [NormedSpace 𝕜 V] [CompleteSpace V noncomputable section open scoped Manifold +/- Next line is necessary while the manifold smoothness class is not extended to `ω`. +Later, replace with `open scoped ContDiff`. -/ +local notation "∞" => (⊤ : ℕ∞) namespace Units diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean b/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean index eb88f2a6c5c87..9386f8ad261ec 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean @@ -27,7 +27,7 @@ charts, differentiable, bijective noncomputable section -open scoped Manifold +open scoped Manifold ContDiff open Bundle Set Topology variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] @@ -99,7 +99,7 @@ theorem mdifferentiableAt_atlas (h : e ∈ atlas H M) {x : M} (hx : x ∈ e.sour ContDiffOn 𝕜 ∞ (I ∘ (chartAt H x).symm.trans e ∘ I.symm) (I.symm ⁻¹' ((chartAt H x).symm.trans e).source ∩ range I) := this.1 - have B := A.differentiableOn le_top (I ((chartAt H x : M → H) x)) mem + have B := A.differentiableOn (mod_cast le_top) (I ((chartAt H x : M → H) x)) mem simp only [mfld_simps] at B rw [inter_comm, differentiableWithinAt_inter] at B · simpa only [mfld_simps] @@ -120,7 +120,7 @@ theorem mdifferentiableAt_atlas_symm (h : e ∈ atlas H M) {x : H} (hx : x ∈ e ContDiffOn 𝕜 ∞ (I ∘ e.symm.trans (chartAt H (e.symm x)) ∘ I.symm) (I.symm ⁻¹' (e.symm.trans (chartAt H (e.symm x))).source ∩ range I) := this.1 - have B := A.differentiableOn le_top (I x) mem + have B := A.differentiableOn (mod_cast le_top) (I x) mem simp only [mfld_simps] at B rw [inter_comm, differentiableWithinAt_inter] at B · simpa only [mfld_simps] diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean index 377ab9d17b795..acf6be63394ec 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean @@ -467,7 +467,8 @@ theorem ContMDiffWithinAt.mdifferentiableWithinAt (hf : ContMDiffWithinAt I I' n apply hf.1.preimage_mem_nhdsWithin exact extChartAt_source_mem_nhds (f x) rw [mdifferentiableWithinAt_iff] - exact ⟨hf.1.mono inter_subset_left, (hf.2.differentiableWithinAt hn).mono (by mfld_set_tac)⟩ + exact ⟨hf.1.mono inter_subset_left, (hf.2.differentiableWithinAt (mod_cast hn)).mono + (by mfld_set_tac)⟩ theorem ContMDiffAt.mdifferentiableAt (hf : ContMDiffAt I I' n f x) (hn : 1 ≤ n) : MDifferentiableAt I I' f x := diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean b/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean index 79091c4619038..b7c3206ec3989 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean @@ -99,7 +99,7 @@ derivative, manifold noncomputable section -open scoped Topology +open scoped Topology ContDiff open Set ChartedSpace section DerivativesDefinitions @@ -147,7 +147,7 @@ theorem differentiableWithinAtProp_self_target {f : H → E'} {s : Set H} {x : H /-- Being differentiable in the model space is a local property, invariant under smooth maps. Therefore, it will lift nicely to manifolds. -/ theorem differentiableWithinAt_localInvariantProp : - (contDiffGroupoid ⊤ I).LocalInvariantProp (contDiffGroupoid ⊤ I') + (contDiffGroupoid ∞ I).LocalInvariantProp (contDiffGroupoid ∞ I') (DifferentiableWithinAtProp I I') := { is_local := by intro s x u f u_open xu @@ -167,7 +167,8 @@ theorem differentiableWithinAt_localInvariantProp : rw [this] at h have : I (e x) ∈ I.symm ⁻¹' e.target ∩ Set.range I := by simp only [hx, mfld_simps] have := (mem_groupoid_of_pregroupoid.2 he).2.contDiffWithinAt this - convert (h.comp' _ (this.differentiableWithinAt le_top)).mono_of_mem_nhdsWithin _ using 1 + convert (h.comp' _ (this.differentiableWithinAt (mod_cast le_top))).mono_of_mem_nhdsWithin _ + using 1 · ext y; simp only [mfld_simps] refine mem_nhdsWithin.mpr @@ -187,7 +188,7 @@ theorem differentiableWithinAt_localInvariantProp : have A : (I' ∘ f ∘ I.symm) (I x) ∈ I'.symm ⁻¹' e'.source ∩ Set.range I' := by simp only [hx, mfld_simps] have := (mem_groupoid_of_pregroupoid.2 he').1.contDiffWithinAt A - convert (this.differentiableWithinAt le_top).comp _ h _ + convert (this.differentiableWithinAt (mod_cast le_top)).comp _ h _ · ext y; simp only [mfld_simps] · intro y hy; simp only [mfld_simps] at hy; simpa only [hy, mfld_simps] using hs hy.1 } diff --git a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean index 3d001e4e01c5d..5a4131d403b71 100644 --- a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean +++ b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean @@ -59,6 +59,9 @@ universe uι uE uH uM uF open Function Filter Module Set open scoped Topology Manifold +/- Next line is necessary while the manifold smoothness class is not extended to `ω`. +Later, replace with `open scoped ContDiff`. -/ +local notation "∞" => (⊤ : ℕ∞) noncomputable section diff --git a/Mathlib/Geometry/Manifold/Sheaf/LocallyRingedSpace.lean b/Mathlib/Geometry/Manifold/Sheaf/LocallyRingedSpace.lean index 48aec60a959fa..5c91ede6dc259 100644 --- a/Mathlib/Geometry/Manifold/Sheaf/LocallyRingedSpace.lean +++ b/Mathlib/Geometry/Manifold/Sheaf/LocallyRingedSpace.lean @@ -31,6 +31,10 @@ smooth manifolds. noncomputable section universe u +/- Next line is necessary while the manifold smoothness class is not extended to `ω`. +Later, replace with `open scoped ContDiff`. -/ +local notation "∞" => (⊤ : ℕ∞) + variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] [CompleteSpace 𝕜] {EM : Type*} [NormedAddCommGroup EM] [NormedSpace 𝕜 EM] {HM : Type*} [TopologicalSpace HM] (IM : ModelWithCorners 𝕜 EM HM) diff --git a/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean b/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean index 0f528359115a1..d794651ff25d8 100644 --- a/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean +++ b/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean @@ -66,6 +66,10 @@ https://github.com/leanprover-community/mathlib4/pull/5726. noncomputable section open TopologicalSpace Opposite +/- Next line is necessary while the manifold smoothness class is not extended to `ω`. +Later, replace with `open scoped ContDiff`. -/ +local notation "∞" => (⊤ : ℕ∞) + universe u variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index df9b1d40d1761..1e7be0aaf95d1 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -133,10 +133,7 @@ universe u v w u' v' w' open Set Filter Function -open scoped Manifold Filter Topology - -/-- The extended natural number `∞` -/ -scoped[Manifold] notation "∞" => (⊤ : ℕ∞) +open scoped Manifold Filter Topology ContDiff /-! ### Models with corners. -/ @@ -529,9 +526,9 @@ section contDiffGroupoid /-! ### Smooth functions on models with corners -/ -variable {m n : ℕ∞} {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] - [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} {M : Type*} - [TopologicalSpace M] +variable {m n : WithTop ℕ∞} {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} + [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] + {I : ModelWithCorners 𝕜 E H} {M : Type*} [TopologicalSpace M] variable (n I) in /-- Given a model with corners `(E, H)`, we define the pregroupoid of `C^n` transformations of `H` @@ -623,8 +620,8 @@ variable {E' H' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] [Topologi /-- The product of two smooth partial homeomorphisms is smooth. -/ theorem contDiffGroupoid_prod {I : ModelWithCorners 𝕜 E H} {I' : ModelWithCorners 𝕜 E' H'} - {e : PartialHomeomorph H H} {e' : PartialHomeomorph H' H'} (he : e ∈ contDiffGroupoid ⊤ I) - (he' : e' ∈ contDiffGroupoid ⊤ I') : e.prod e' ∈ contDiffGroupoid ⊤ (I.prod I') := by + {e : PartialHomeomorph H H} {e' : PartialHomeomorph H' H'} (he : e ∈ contDiffGroupoid ∞ I) + (he' : e' ∈ contDiffGroupoid ∞ I') : e.prod e' ∈ contDiffGroupoid ∞ (I.prod I') := by cases' he with he he_symm cases' he' with he' he'_symm simp only at he he_symm he' he'_symm @@ -672,7 +669,7 @@ theorem smoothManifoldWithCorners_of_contDiffOn {𝕜 : Type*} [NontriviallyNorm {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) (M : Type*) [TopologicalSpace M] [ChartedSpace H M] (h : ∀ e e' : PartialHomeomorph M H, e ∈ atlas H M → e' ∈ atlas H M → - ContDiffOn 𝕜 ⊤ (I ∘ e.symm ≫ₕ e' ∘ I.symm) (I.symm ⁻¹' (e.symm ≫ₕ e').source ∩ range I)) : + ContDiffOn 𝕜 ∞ (I ∘ e.symm ≫ₕ e' ∘ I.symm) (I.symm ⁻¹' (e.symm ≫ₕ e').source ∩ range I)) : SmoothManifoldWithCorners I M where compatible := by haveI : HasGroupoid M (contDiffGroupoid ∞ I) := hasGroupoid_of_pregroupoid _ (h _ _) @@ -741,8 +738,8 @@ instance prod {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedA compatible := by rintro f g ⟨f1, hf1, f2, hf2, rfl⟩ ⟨g1, hg1, g2, hg2, rfl⟩ rw [PartialHomeomorph.prod_symm, PartialHomeomorph.prod_trans] - have h1 := (contDiffGroupoid ⊤ I).compatible hf1 hg1 - have h2 := (contDiffGroupoid ⊤ I').compatible hf2 hg2 + have h1 := (contDiffGroupoid ∞ I).compatible hf1 hg1 + have h2 := (contDiffGroupoid ∞ I').compatible hf2 hg2 exact contDiffGroupoid_prod h1 h2 end SmoothManifoldWithCorners @@ -1065,13 +1062,13 @@ open SmoothManifoldWithCorners theorem contDiffOn_extend_coord_change [ChartedSpace H M] (hf : f ∈ maximalAtlas I M) (hf' : f' ∈ maximalAtlas I M) : - ContDiffOn 𝕜 ⊤ (f.extend I ∘ (f'.extend I).symm) ((f'.extend I).symm ≫ f.extend I).source := by + ContDiffOn 𝕜 ∞ (f.extend I ∘ (f'.extend I).symm) ((f'.extend I).symm ≫ f.extend I).source := by rw [extend_coord_change_source, I.image_eq] exact (StructureGroupoid.compatible_of_mem_maximalAtlas hf' hf).1 theorem contDiffWithinAt_extend_coord_change [ChartedSpace H M] (hf : f ∈ maximalAtlas I M) (hf' : f' ∈ maximalAtlas I M) {x : E} (hx : x ∈ ((f'.extend I).symm ≫ f.extend I).source) : - ContDiffWithinAt 𝕜 ⊤ (f.extend I ∘ (f'.extend I).symm) (range I) x := by + ContDiffWithinAt 𝕜 ∞ (f.extend I ∘ (f'.extend I).symm) (range I) x := by apply (contDiffOn_extend_coord_change hf hf' x hx).mono_of_mem_nhdsWithin rw [extend_coord_change_source] at hx ⊢ obtain ⟨z, hz, rfl⟩ := hx @@ -1079,7 +1076,7 @@ theorem contDiffWithinAt_extend_coord_change [ChartedSpace H M] (hf : f ∈ maxi theorem contDiffWithinAt_extend_coord_change' [ChartedSpace H M] (hf : f ∈ maximalAtlas I M) (hf' : f' ∈ maximalAtlas I M) {x : M} (hxf : x ∈ f.source) (hxf' : x ∈ f'.source) : - ContDiffWithinAt 𝕜 ⊤ (f.extend I ∘ (f'.extend I).symm) (range I) (f'.extend I x) := by + ContDiffWithinAt 𝕜 ∞ (f.extend I ∘ (f'.extend I).symm) (range I) (f'.extend I x) := by refine contDiffWithinAt_extend_coord_change hf hf' ?_ rw [← extend_image_source_inter] exact mem_image_of_mem _ ⟨hxf', hxf⟩ @@ -1408,15 +1405,14 @@ theorem ext_coord_change_source (x x' : M) : open SmoothManifoldWithCorners theorem contDiffOn_ext_coord_change [SmoothManifoldWithCorners I M] (x x' : M) : - ContDiffOn 𝕜 ⊤ (extChartAt I x ∘ (extChartAt I x').symm) + ContDiffOn 𝕜 ∞ (extChartAt I x ∘ (extChartAt I x').symm) ((extChartAt I x').symm ≫ extChartAt I x).source := contDiffOn_extend_coord_change (chart_mem_maximalAtlas x) (chart_mem_maximalAtlas x') theorem contDiffWithinAt_ext_coord_change [SmoothManifoldWithCorners I M] (x x' : M) {y : E} (hy : y ∈ ((extChartAt I x').symm ≫ extChartAt I x).source) : - ContDiffWithinAt 𝕜 ⊤ (extChartAt I x ∘ (extChartAt I x').symm) (range I) y := - contDiffWithinAt_extend_coord_change (chart_mem_maximalAtlas x) (chart_mem_maximalAtlas x') - hy + ContDiffWithinAt 𝕜 ∞ (extChartAt I x ∘ (extChartAt I x').symm) (range I) y := + contDiffWithinAt_extend_coord_change (chart_mem_maximalAtlas x) (chart_mem_maximalAtlas x') hy variable (I I') in /-- Conjugating a function to write it in the preferred charts around `x`. diff --git a/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean b/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean index 751bf15036342..60a27aee36286 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean @@ -19,6 +19,9 @@ sections of a smooth vector bundle over a manifold `M` and prove that it's a mod open Bundle Filter Function open scoped Bundle Manifold +/- Next line is necessary while the manifold smoothness class is not extended to `ω`. +Later, replace with `open scoped ContDiff`. -/ +local notation "∞" => (⊤ : ℕ∞) variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean index f40f1f950c07a..5931fefe4f64b 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean @@ -36,7 +36,7 @@ and a smooth manifold. open Bundle Set SmoothManifoldWithCorners PartialHomeomorph ContinuousLinearMap -open scoped Manifold Topology Bundle +open scoped Manifold Topology Bundle ContDiff noncomputable section @@ -49,7 +49,6 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCom [SmoothManifoldWithCorners I M] {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] [SmoothManifoldWithCorners I' M'] {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] - /-- Auxiliary lemma for tangent spaces: the derivative of a coordinate change between two charts is smooth on its source. -/ theorem contDiffOn_fderiv_coord_change (i j : atlas H M) : @@ -58,12 +57,12 @@ 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.uniqueDiffOn le_top <| h hx).mono h + refine (ContDiffWithinAt.fderivWithin_right ?_ I.uniqueDiffOn (n := ∞) (mod_cast le_top) + <| h hx).mono h refine (PartialHomeomorph.contDiffOn_extend_coord_change (subset_maximalAtlas j.2) (subset_maximalAtlas i.2) x hx).mono_of_mem_nhdsWithin ?_ exact i.1.extend_coord_change_source_mem_nhdsWithin j.1 hx - open SmoothManifoldWithCorners variable (I M) in @@ -104,9 +103,9 @@ def tangentBundleCore : VectorBundleCore 𝕜 M E (atlas H M) where simp_rw [Function.comp_apply, (j.1.extend I).left_inv hy] · simp_rw [Function.comp_apply, i.1.extend_left_inv hxi, j.1.extend_left_inv hxj] · exact (contDiffWithinAt_extend_coord_change' (subset_maximalAtlas k.2) - (subset_maximalAtlas j.2) hxk hxj).differentiableWithinAt le_top + (subset_maximalAtlas j.2) hxk hxj).differentiableWithinAt (mod_cast le_top) · exact (contDiffWithinAt_extend_coord_change' (subset_maximalAtlas j.2) - (subset_maximalAtlas i.2) hxj hxi).differentiableWithinAt le_top + (subset_maximalAtlas i.2) hxj hxi).differentiableWithinAt (mod_cast le_top) · intro x _; exact mem_range_self _ · exact I.uniqueDiffWithinAt_image · rw [Function.comp_apply, i.1.extend_left_inv hxi] diff --git a/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean b/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean index 5551413ea3416..a6416ab861ef1 100644 --- a/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean +++ b/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean @@ -33,6 +33,9 @@ variable {ι : Type uι} {E : Type uE} [NormedAddCommGroup E] [NormedSpace ℝ E open Function Filter Module Set Topology open scoped Manifold +/- Next line is necessary while the manifold smoothness class is not extended to `ω`. +Later, replace with `open scoped ContDiff`. -/ +local notation "∞" => (⊤ : ℕ∞) noncomputable section diff --git a/MathlibTest/fun_prop2.lean b/MathlibTest/fun_prop2.lean index e98b764c72c0a..8a5d906d3f74b 100644 --- a/MathlibTest/fun_prop2.lean +++ b/MathlibTest/fun_prop2.lean @@ -76,7 +76,7 @@ example : AEMeasurable T := by fun_prop -private theorem t1 : (5: ℕ) + (1 : ℕ∞) ≤ (12 : ℕ∞) := by norm_cast +private theorem t1 : (5: ℕ) + (1 : ℕ∞) ≤ (12 : WithTop ℕ∞) := by norm_cast example {f : ℝ → ℝ} (hf : ContDiff ℝ 12 f) : Differentiable ℝ (iteratedDeriv 5 (fun x => f (2*(f (x + x))) + x)) := by From beb3168d9638a2a180f3b33e673e3ac69444c965 Mon Sep 17 00:00:00 2001 From: Antoine Chambert-Loir Date: Mon, 25 Nov 2024 12:42:00 +0000 Subject: [PATCH 285/829] feat: additivize equivariance of morphisms of actions (#19171) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Additivize the definition of equivariant morphisms with respect to multiplicative actions. The notation introduced is `X →ₑᵥ[φ] Y` for equivariant morphisms, and `X →ᵥ[M] Y` when the underlying map is identity. The goal is to be able to fully additivize the properties of blocks for actions and the definition of primitive actions in #12052 Co-authored-by: leanprover-community-mathlib4-bot --- Mathlib/GroupTheory/GroupAction/Hom.lean | 101 ++++++++++++++++++----- 1 file changed, 80 insertions(+), 21 deletions(-) diff --git a/Mathlib/GroupTheory/GroupAction/Hom.lean b/Mathlib/GroupTheory/GroupAction/Hom.lean index 95bd4983abab0..2ce20e3e8812b 100644 --- a/Mathlib/GroupTheory/GroupAction/Hom.lean +++ b/Mathlib/GroupTheory/GroupAction/Hom.lean @@ -15,6 +15,7 @@ import Mathlib.Algebra.Group.Hom.CompTypeclasses * `MulActionHom φ X Y`, the type of equivariant functions from `X` to `Y`, where `φ : M → N` is a map, `M` acting on the type `X` and `N` acting on the type of `Y`. + `AddActionHom φ X Y` is its additive version. * `DistribMulActionHom φ A B`, the type of equivariant additive monoid homomorphisms from `A` to `B`, where `φ : M → N` is a morphism of monoids, @@ -25,7 +26,8 @@ import Mathlib.Algebra.Group.Hom.CompTypeclasses The above types have corresponding classes: * `MulActionHomClass F φ X Y` states that `F` is a type of bundled `X → Y` homs - which are `φ`-equivariant + which are `φ`-equivariant; + `AddActionHomClass F φ X Y` is its additive version. * `DistribMulActionHomClass F φ A B` states that `F` is a type of bundled `A → B` homs preserving the additive monoid structure and `φ`-equivariant * `SMulSemiringHomClass F φ R S` states that `F` is a type of bundled `R → S` homs @@ -35,14 +37,19 @@ The above types have corresponding classes: We introduce the following notation to code equivariant maps (the subscript index `ₑ` is for *equivariant*) : -* `X →ₑ[φ] Y` is `MulActionHom φ X Y`. +* `X →ₑ[φ] Y` is `MulActionHom φ X Y` and `AddActionHom φ X Y` * `A →ₑ+[φ] B` is `DistribMulActionHom φ A B`. * `R →ₑ+*[φ] S` is `MulSemiringActionHom φ R S`. When `M = N` and `φ = MonoidHom.id M`, we provide the backward compatible notation : -* `X →[M] Y` is `MulActionHom (@id M) X Y` +* `X →[M] Y` is `MulActionHom (@id M) X Y` and `AddActionHom (@id M) X Y` * `A →+[M] B` is `DistribMulActionHom (MonoidHom.id M) A B` * `R →+*[M] S` is `MulSemiringActionHom (MonoidHom.id M) R S` + +The notation for `MulActionHom` and `AddActionHom` is the same, because it is unlikely +that it could lead to confusion — unless one needs types `M` and `X` with simultaneous +instances of `Mul M`, `Add M`, `SMul M X` and `VAdd M X`… + -/ assert_not_exists Submonoid @@ -56,11 +63,21 @@ variable (X : Type*) [SMul M X] [SMul M' X] variable (Y : Type*) [SMul N Y] [SMul M' Y] variable (Z : Type*) [SMul P Z] +/-- Equivariant functions : +When `φ : M → N` is a function, and types `X` and `Y` are endowed with additive actions +of `M` and `N`, a function `f : X → Y` is `φ`-equivariant if `f (m +ᵥ x) = (φ m) +ᵥ (f x)`. -/ +structure AddActionHom {M N : Type*} (φ: M → N) (X : Type*) [VAdd M X] (Y : Type*) [VAdd N Y] where + /-- The underlying function. -/ + protected toFun : X → Y + /-- The proposition that the function commutes with the additive actions. -/ + protected map_vadd' : ∀ (m : M) (x : X), toFun (m +ᵥ x) = (φ m) +ᵥ toFun x + /-- Equivariant functions : When `φ : M → N` is a function, and types `X` and `Y` are endowed with actions of `M` and `N`, a function `f : X → Y` is `φ`-equivariant if `f (m • x) = (φ m) • (f x)`. -/ -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. -- @[nolint has_nonempty_instance] +@[to_additive] structure MulActionHom where /-- The underlying function. -/ protected toFun : X → Y @@ -70,20 +87,40 @@ structure MulActionHom where /- Porting note: local notation given a name, conflict with Algebra.Hom.GroupAction see https://github.com/leanprover/lean4/issues/2000 -/ /-- `φ`-equivariant functions `X → Y`, -where `φ : M → N`, where `M` and `N` act on `X` and `Y` respectively -/ +where `φ : M → N`, where `M` and `N` act on `X` and `Y` respectively.-/ notation:25 (name := «MulActionHomLocal≺») X " →ₑ[" φ:25 "] " Y:0 => MulActionHom φ X Y -/-- `M`-equivariant functions `X → Y` with respect to the action of `M` - -This is the same as `X →ₑ[@id M] Y` -/ +/-- `M`-equivariant functions `X → Y` with respect to the action of `M`. +This is the same as `X →ₑ[@id M] Y`. -/ notation:25 (name := «MulActionHomIdLocal≺») X " →[" M:25 "] " Y:0 => MulActionHom (@id M) X Y +/-- `φ`-equivariant functions `X → Y`, +where `φ : M → N`, where `M` and `N` act additively on `X` and `Y` respectively +We use the same notation as for multiplicative actions, as conflicts are unlikely. -/ +notation:25 (name := «AddActionHomLocal≺») X " →ₑ[" φ:25 "] " Y:0 => AddActionHom φ X Y + +/-- `M`-equivariant functions `X → Y` with respect to the additive action of `M`. +This is the same as `X →ₑ[@id M] Y`. + +We use the same notation as for multiplicative actions, as conflicts are unlikely. -/ +notation:25 (name := «AddActionHomIdLocal≺») X " →[" M:25 "] " Y:0 => AddActionHom (@id M) X Y + +/-- `AddActionSemiHomClass F φ X Y` states that + `F` is a type of morphisms which are `φ`-equivariant. + +You should extend this class when you extend `AddActionHom`. -/ +class AddActionSemiHomClass (F : Type*) + {M N : outParam Type*} (φ : outParam (M → N)) + (X Y : outParam Type*) [VAdd M X] [VAdd N Y] [FunLike F X Y] : Prop where + /-- The proposition that the function preserves the action. -/ + map_vaddₛₗ : ∀ (f : F) (c : M) (x : X), f (c +ᵥ x) = (φ c) +ᵥ (f x) /-- `MulActionSemiHomClass F φ X Y` states that `F` is a type of morphisms which are `φ`-equivariant. You should extend this class when you extend `MulActionHom`. -/ +@[to_additive] class MulActionSemiHomClass (F : Type*) {M N : outParam Type*} (φ : outParam (M → N)) (X Y : outParam Type*) [SMul M X] [SMul N Y] [FunLike F X Y] : Prop where @@ -91,19 +128,23 @@ class MulActionSemiHomClass (F : Type*) map_smulₛₗ : ∀ (f : F) (c : M) (x : X), f (c • x) = (φ c) • (f x) export MulActionSemiHomClass (map_smulₛₗ) +export AddActionSemiHomClass (map_vaddₛₗ) /-- `MulActionHomClass F M X Y` states that `F` is a type of morphisms which are equivariant with respect to actions of `M` This is an abbreviation of `MulActionSemiHomClass`. -/ +@[to_additive "`MulActionHomClass F M X Y` states that `F` is a type of +morphisms which are equivariant with respect to actions of `M` +This is an abbreviation of `MulActionSemiHomClass`."] abbrev MulActionHomClass (F : Type*) (M : outParam Type*) (X Y : outParam Type*) [SMul M X] [SMul M Y] [FunLike F X Y] := MulActionSemiHomClass F (@id M) X Y -instance : FunLike (MulActionHom φ X Y) X Y where +@[to_additive] instance : FunLike (MulActionHom φ X Y) X Y where coe := MulActionHom.toFun coe_injective' f g h := by cases f; cases g; congr -@[simp] +@[to_additive (attr := simp)] theorem map_smul {F M X Y : Type*} [SMul M X] [SMul M Y] [FunLike F X Y] [MulActionHomClass F M X Y] (f : F) (c : M) (x : X) : f (c • x) = c • f x := @@ -113,10 +154,12 @@ theorem map_smul {F M X Y : Type*} [SMul M X] [SMul M Y] -- Porting note: removed has_coe_to_fun instance, coercions handled differently now +@[to_additive] instance : MulActionSemiHomClass (X →ₑ[φ] Y) φ X Y where map_smulₛₗ := MulActionHom.map_smul' initialize_simps_projections MulActionHom (toFun → apply) +initialize_simps_projections AddActionHom (toFun → apply) namespace MulActionHom @@ -128,7 +171,10 @@ see also Algebra.Hom.Group -/ /-- Turn an element of a type `F` satisfying `MulActionSemiHomClass F φ X Y` into an actual `MulActionHom`. This is declared as the default coercion from `F` to `MulActionSemiHom φ X Y`. -/ -@[coe] +@[to_additive (attr := coe) + "Turn an element of a type `F` satisfying `AddActionSemiHomClass F φ X Y` + into an actual `AddActionHom`. + This is declared as the default coercion from `F` to `AddActionSemiHom φ X Y`."] def _root_.MulActionSemiHomClass.toMulActionHom [MulActionSemiHomClass F φ X Y] (f : F) : X →ₑ[φ] Y where toFun := DFunLike.coe f @@ -136,39 +182,44 @@ def _root_.MulActionSemiHomClass.toMulActionHom [MulActionSemiHomClass F φ X Y] /-- Any type satisfying `MulActionSemiHomClass` can be cast into `MulActionHom` via `MulActionHomSemiClass.toMulActionHom`. -/ +@[to_additive] instance [MulActionSemiHomClass F φ X Y] : CoeTC F (X →ₑ[φ] Y) := ⟨MulActionSemiHomClass.toMulActionHom⟩ variable (M' X Y F) in /-- If Y/X/M forms a scalar tower, any map X → Y preserving X-action also preserves M-action. -/ +@[to_additive] theorem _root_.IsScalarTower.smulHomClass [MulOneClass X] [SMul X Y] [IsScalarTower M' X Y] [MulActionHomClass F X X Y] : MulActionHomClass F M' X Y where map_smulₛₗ f m x := by rw [← mul_one (m • x), ← smul_eq_mul, map_smul, smul_assoc, ← map_smul, smul_eq_mul, mul_one, id_eq] +@[to_additive] protected theorem map_smul (f : X →[M'] Y) (m : M') (x : X) : f (m • x) = m • f x := map_smul f m x -@[ext] +@[to_additive (attr := ext)] theorem ext {f g : X →ₑ[φ] Y} : (∀ x, f x = g x) → f = g := DFunLike.ext f g +@[to_additive] protected theorem congr_fun {f g : X →ₑ[φ] Y} (h : f = g) (x : X) : f x = g x := DFunLike.congr_fun h _ /-- Two equal maps on scalars give rise to an equivariant map for identity -/ +@[to_additive "Two equal maps on scalars give rise to an equivariant map for identity"] def ofEq {φ' : M → N} (h : φ = φ') (f : X →ₑ[φ] Y) : X →ₑ[φ'] Y where toFun := f.toFun map_smul' m a := h ▸ f.map_smul' m a -@[simp] +@[to_additive (attr := simp)] theorem ofEq_coe {φ' : M → N} (h : φ = φ') (f : X →ₑ[φ] Y) : (f.ofEq h).toFun = f.toFun := rfl -@[simp] +@[to_additive (attr := simp)] theorem ofEq_apply {φ' : M → N} (h : φ = φ') (f : X →ₑ[φ] Y) (a : X) : (f.ofEq h) a = f a := rfl @@ -177,12 +228,13 @@ theorem ofEq_apply {φ' : M → N} (h : φ = φ') (f : X →ₑ[φ] Y) (a : X) : variable {ψ χ} (M N) /-- The identity map as an equivariant map. -/ +@[to_additive "The identity map as an equivariant map."] protected def id : X →[M] X := ⟨id, fun _ _ => rfl⟩ variable {M N Z} -@[simp] +@[to_additive (attr := simp)] theorem id_apply (x : X) : MulActionHom.id M x = x := rfl @@ -197,6 +249,7 @@ variable {φ ψ χ X Y Z} -- attribute [instance] CompTriple.id_comp CompTriple.comp_id /-- Composition of two equivariant maps. -/ +@[to_additive "Composition of two equivariant maps."] def comp (g : Y →ₑ[ψ] Z) (f : X →ₑ[φ] Y) [κ : CompTriple φ ψ χ] : X →ₑ[χ] Z := ⟨g ∘ f, fun m x => @@ -206,22 +259,22 @@ def comp (g : Y →ₑ[ψ] Z) (f : X →ₑ[φ] Y) [κ : CompTriple φ ψ χ] : _ = (ψ ∘ φ) m • g (f x) := rfl _ = χ m • g (f x) := by rw [κ.comp_eq] ⟩ -@[simp] +@[to_additive (attr := simp)] theorem comp_apply (g : Y →ₑ[ψ] Z) (f : X →ₑ[φ] Y) [CompTriple φ ψ χ] (x : X) : g.comp f x = g (f x) := rfl -@[simp] +@[to_additive (attr := simp)] theorem id_comp (f : X →ₑ[φ] Y) : (MulActionHom.id N).comp f = f := ext fun x => by rw [comp_apply, id_apply] -@[simp] +@[to_additive (attr := simp)] theorem comp_id (f : X →ₑ[φ] Y) : f.comp (MulActionHom.id M) = f := ext fun x => by rw [comp_apply, id_apply] -@[simp] +@[to_additive (attr := simp)] theorem comp_assoc {Q T : Type*} [SMul Q T] {η : P → Q} {θ : M → Q} {ζ : N → Q} (h : Z →ₑ[η] T) (g : Y →ₑ[ψ] Z) (f : X →ₑ[φ] Y) @@ -232,8 +285,9 @@ theorem comp_assoc {Q T : Type*} [SMul Q T] variable {φ' : N → M} variable {Y₁ : Type*} [SMul M Y₁] + /-- The inverse of a bijective equivariant map is equivariant. -/ -@[simps] +@[to_additive (attr := simps) "The inverse of a bijective equivariant map is equivariant."] def inverse (f : X →[M] Y₁) (g : Y₁ → X) (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) : Y₁ →[M] X where toFun := g @@ -245,7 +299,7 @@ def inverse (f : X →[M] Y₁) (g : Y₁ → X) /-- The inverse of a bijective equivariant map is equivariant. -/ -@[simps] +@[to_additive (attr := simps) "The inverse of a bijective equivariant map is equivariant."] def inverse' (f : X →ₑ[φ] Y) (g : Y → X) (k : Function.RightInverse φ' φ) (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) : Y →ₑ[φ'] X where @@ -257,11 +311,13 @@ def inverse' (f : X →ₑ[φ] Y) (g : Y → X) (k : Function.RightInverse φ' _ = g (f (φ' m • g x)) := by rw [map_smulₛₗ] _ = φ' m • g x := by rw [h₁] +@[to_additive] lemma inverse_eq_inverse' (f : X →[M] Y₁) (g : Y₁ → X) (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) : inverse f g h₁ h₂ = inverse' f g (congrFun rfl) h₁ h₂ := by rfl +@[to_additive] theorem inverse'_inverse' {f : X →ₑ[φ] Y} {g : Y → X} {k₁ : Function.LeftInverse φ' φ} {k₂ : Function.RightInverse φ' φ} @@ -269,6 +325,7 @@ theorem inverse'_inverse' inverse' (inverse' f g k₂ h₁ h₂) f k₁ h₂ h₁ = f := ext fun _ => rfl +@[to_additive] theorem comp_inverse' {f : X →ₑ[φ] Y} {g : Y → X} {k₁ : Function.LeftInverse φ' φ} {k₂ : Function.RightInverse φ' φ} {h₁ : Function.LeftInverse g f} {h₂ : Function.RightInverse g f} : @@ -279,6 +336,7 @@ theorem comp_inverse' {f : X →ₑ[φ] Y} {g : Y → X} simp only [comp_apply, inverse_apply, id_apply] exact h₁ x +@[to_additive] theorem inverse'_comp {f : X →ₑ[φ] Y} {g : Y → X} {k₂ : Function.RightInverse φ' φ} {h₁ : Function.LeftInverse g f} {h₂ : Function.RightInverse g f} : @@ -290,7 +348,8 @@ theorem inverse'_comp {f : X →ₑ[φ] Y} {g : Y → X} /-- If actions of `M` and `N` on `α` commute, then for `c : M`, `(c • · : α → α)` is an `N`-action homomorphism. -/ -@[simps] +@[to_additive (attr := simps) "If additive actions of `M` and `N` on `α` commute, + then for `c : M`, `(c • · : α → α)` is an `N`-additive action homomorphism."] def _root_.SMulCommClass.toMulActionHom {M} (N α : Type*) [SMul M α] [SMul N α] [SMulCommClass M N α] (c : M) : α →[N] α where From 7f28d7f73f74b31f72449bb9eb240db3ca68fe38 Mon Sep 17 00:00:00 2001 From: Antoine Chambert-Loir Date: Mon, 25 Nov 2024 12:42:01 +0000 Subject: [PATCH 286/829] feat(Data/Finsupp/MonomialOrder): monomial orders (#19453) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Definition of monomial orders, with the example of the lexicographic ordering This is the starting point of implementation of the beginnings of Gröbner theory. 3 subsequent PRs will address: - division algorithm - example of the homogeneous lexicographic ordering - example of the homogeneous reverse lexicographic ordering These various orderings (and other ones…) may look exotic but are actually used in algebraic geometry. --- Mathlib.lean | 1 + Mathlib/Data/Finsupp/MonomialOrder.lean | 156 ++++++++++++++++++++++++ docs/references.bib | 15 +++ 3 files changed, 172 insertions(+) create mode 100644 Mathlib/Data/Finsupp/MonomialOrder.lean diff --git a/Mathlib.lean b/Mathlib.lean index 34090017c1344..f73b287687527 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2449,6 +2449,7 @@ import Mathlib.Data.Finsupp.Fintype import Mathlib.Data.Finsupp.Indicator import Mathlib.Data.Finsupp.Interval import Mathlib.Data.Finsupp.Lex +import Mathlib.Data.Finsupp.MonomialOrder import Mathlib.Data.Finsupp.Multiset import Mathlib.Data.Finsupp.NeLocus import Mathlib.Data.Finsupp.Notation diff --git a/Mathlib/Data/Finsupp/MonomialOrder.lean b/Mathlib/Data/Finsupp/MonomialOrder.lean new file mode 100644 index 0000000000000..431fcb7d3ce34 --- /dev/null +++ b/Mathlib/Data/Finsupp/MonomialOrder.lean @@ -0,0 +1,156 @@ +/- +Copyright (c) 2024 Antoine Chambert-Loir. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Antoine Chambert-Loir +-/ +import Mathlib.Data.Finsupp.Lex +import Mathlib.Data.Finsupp.WellFounded +import Mathlib.Data.List.TFAE + +/-! # Monomial orders + +## Monomial orders + +A *monomial order* is well ordering relation on a type of the form `σ →₀ ℕ` which +is compatible with addition and for which `0` is the smallest element. +Since several monomial orders may have to be used simultaneously, one cannot +get them as instances. +In this formalization, they are presented as a structure `MonomialOrder` which encapsulates +`MonomialOrder.toSyn`, an additive and monotone isomorphism to a linearly ordered cancellative +additive commutative monoid. +The entry `MonomialOrder.wf` asserts that `MonomialOrder.syn` is well founded. + +The terminology comes from commutative algebra and algebraic geometry, especially Gröbner bases, +where `c : σ →₀ ℕ` are exponents of monomials. + +Given a monomial order `m : MonomialOrder σ`, we provide the notation +`c ≼[m] d` and `c ≺[m] d` to compare `c d : σ →₀ ℕ` with respect to `m`. +It is activated using `open scoped MonomialOrder`. + +## Examples + +Commutative algebra defines many monomial orders, with different usefulness ranges. +In this file, we provide the basic example of lexicographic ordering. +For the graded lexicographic ordering, see `Mathlib/Data/Finsupp/DegLex.lean` + +* `MonomialOrder.lex` : the lexicographic ordering on `σ →₀ ℕ`. +For this, `σ` needs to be embedded with an ordering relation which satisfies `WellFoundedGT σ`. +(This last property is automatic when `σ` is finite). + +The type synonym is `Lex (σ →₀ ℕ)` and the two lemmas `MonomialOrder.lex_le_iff` +and `MonomialOrder.lex_lt_iff` rewrite the ordering as comparisons in the type `Lex (σ →₀ ℕ)`. + +## References + +* [Cox, Little and O'Shea, *Ideals, varieties, and algorithms*][coxlittleoshea1997] +* [Becker and Weispfenning, *Gröbner bases*][Becker-Weispfenning1993] + +## Note + +In algebraic geometry, when the finitely many variables are indexed by integers, +it is customary to order them using the opposite order : `MvPolynomial.X 0 > MvPolynomial.X 1 > … ` + +-/ + +/-- Monomial orders : equivalence of `σ →₀ ℕ` with a well ordered type -/ +structure MonomialOrder (σ : Type*) where + /-- The synonym type -/ + syn : Type* + /-- `syn` is a linearly ordered cancellative additive commutative monoid -/ + locacm : LinearOrderedCancelAddCommMonoid syn := by infer_instance + /-- the additive equivalence from `σ →₀ ℕ` to `syn` -/ + toSyn : (σ →₀ ℕ) ≃+ syn + /-- `toSyn` is monotone -/ + toSyn_monotone : Monotone toSyn + /-- `syn` is a well ordering -/ + wf : WellFoundedLT syn := by infer_instance + +attribute [instance] MonomialOrder.locacm MonomialOrder.wf + +namespace MonomialOrder + +variable {σ : Type*} (m : MonomialOrder σ) + +lemma le_add_right (a b : σ →₀ ℕ) : + m.toSyn a ≤ m.toSyn a + m.toSyn b := by + rw [← map_add] + exact m.toSyn_monotone le_self_add + +instance orderBot : OrderBot (m.syn) where + bot := 0 + bot_le a := by + have := m.le_add_right 0 (m.toSyn.symm a) + simp [map_add, zero_add] at this + exact this + +@[simp] +theorem bot_eq_zero : (⊥ : m.syn) = 0 := rfl + +theorem eq_zero_iff {a : m.syn} : a = 0 ↔ a ≤ 0 := eq_bot_iff + +lemma toSyn_strictMono : StrictMono (m.toSyn) := by + apply m.toSyn_monotone.strictMono_of_injective m.toSyn.injective + +/-- Given a monomial order, notation for the corresponding strict order relation on `σ →₀ ℕ` -/ +scoped +notation:25 c "≺[" m:25 "]" d:25 => (MonomialOrder.toSyn m c < MonomialOrder.toSyn m d) + +/-- Given a monomial order, notation for the corresponding order relation on `σ →₀ ℕ` -/ +scoped +notation:25 c "≼[" m:25 "]" d:25 => (MonomialOrder.toSyn m c ≤ MonomialOrder.toSyn m d) + +end MonomialOrder + +section Lex + +open Finsupp + +open scoped MonomialOrder + +-- The linear order on `Finsupp`s obtained by the lexicographic ordering. -/ +noncomputable instance {α N : Type*} [LinearOrder α] [OrderedCancelAddCommMonoid N] : + OrderedCancelAddCommMonoid (Lex (α →₀ N)) where + le_of_add_le_add_left a b c h := by simpa only [add_le_add_iff_left] using h + add_le_add_left a b h c := by simpa only [add_le_add_iff_left] using h + +theorem Finsupp.lex_lt_iff {α N : Type*} [LinearOrder α] [LinearOrder N] [Zero N] + {a b : Lex (α →₀ N)} : + a < b ↔ ∃ i, (∀ j, j< i → ofLex a j = ofLex b j) ∧ ofLex a i < ofLex b i := + Finsupp.lex_def + +theorem Finsupp.lex_le_iff {α N : Type*} [LinearOrder α] [LinearOrder N] [Zero N] + {a b : Lex (α →₀ N)} : + a ≤ b ↔ a = b ∨ ∃ i, (∀ j, j< i → ofLex a j = ofLex b j) ∧ ofLex a i < ofLex b i := by + rw [le_iff_eq_or_lt, Finsupp.lex_lt_iff] + +/-- for the lexicographic ordering, X 0 * X 1 < X 0 ^ 2 -/ +example : toLex (Finsupp.single 0 2) > toLex (Finsupp.single 0 1 + Finsupp.single 1 1) := by + use 0; simp + +/-- for the lexicographic ordering, X 1 < X 0 -/ +example : toLex (Finsupp.single 1 1) < toLex (Finsupp.single 0 1) := by + use 0; simp + +/-- for the lexicographic ordering, X 1 < X 0 ^ 2 -/ +example : toLex (Finsupp.single 1 1) < toLex (Finsupp.single 0 2) := by + use 0; simp + + +variable {σ : Type*} [LinearOrder σ] + +/-- The lexicographic order on `σ →₀ ℕ`, as a `MonomialOrder` -/ +noncomputable def MonomialOrder.lex [WellFoundedGT σ] : + MonomialOrder σ where + syn := Lex (σ →₀ ℕ) + toSyn := + { toEquiv := toLex + map_add' := toLex_add } + toSyn_monotone := Finsupp.toLex_monotone + +theorem MonomialOrder.lex_le_iff [WellFoundedGT σ] {c d : σ →₀ ℕ} : + c ≼[lex] d ↔ toLex c ≤ toLex d := Iff.rfl + +theorem MonomialOrder.lex_lt_iff [WellFoundedGT σ] {c d : σ →₀ ℕ} : + c ≺[lex] d ↔ toLex c < toLex d := Iff.rfl + +end Lex diff --git a/docs/references.bib b/docs/references.bib index 3d890cec41f2a..0e5ce6b9bd094 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -247,6 +247,21 @@ @Book{ beals2004 year = {2004} } +@Book{ Becker-Weispfenning1993, + address = {New York, NY}, + series = {Graduate Texts in Mathematics}, + title = {Gröbner Bases}, + volume = {141}, + isbn = {978-1-4612-6944-1}, + url = {http://link.springer.com/10.1007/978-1-4612-0913-3}, + doi = {10.1007/978-1-4612-0913-3}, + publisher = {Springer New York}, + author = {Becker, Thomas and Weispfenning, Volker}, + year = {1993}, + collection = {Graduate Texts in Mathematics}, + language = {en} +} + @Book{ behrends1979, author = {Ehrhard {Behrends}}, title = {{M-structure and the Banach-Stone theorem}}, From f26c0a21e7f81a4e95e0703e492af2d729da839d Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Mon, 25 Nov 2024 13:12:47 +0000 Subject: [PATCH 287/829] chore(RingTheory/Noetherian): further split `Defs.lean` (#19364) Now that we have split `RingTheory.Finiteness`, we have the ability to clean up the import hierarchy of `RingTheory.PrincipalIdealDomain` further by splitting `Noetherian/Defs.lean` further. Now `Noetherian/Defs.lean` only contains results relating finite generation of submodule and wellfoundedness of the inclusion relation. --- Mathlib.lean | 1 + .../Algebra/Module/FinitePresentation.lean | 2 +- Mathlib/FieldTheory/Tower.lean | 2 +- .../RingTheory/GradedAlgebra/Noetherian.lean | 2 +- .../RingTheory/Ideal/Quotient/Noetherian.lean | 2 +- Mathlib/RingTheory/Noetherian/Basic.lean | 366 ++++++++++++++++++ Mathlib/RingTheory/Noetherian/Defs.lean | 298 +------------- Mathlib/RingTheory/Noetherian/Orzech.lean | 2 +- Mathlib/RingTheory/Polynomial/Basic.lean | 2 +- Mathlib/RingTheory/PrimeSpectrum.lean | 2 +- Mathlib/RingTheory/PrincipalIdealDomain.lean | 1 + 11 files changed, 379 insertions(+), 301 deletions(-) create mode 100644 Mathlib/RingTheory/Noetherian/Basic.lean diff --git a/Mathlib.lean b/Mathlib.lean index f73b287687527..39dccdce392ff 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4385,6 +4385,7 @@ import Mathlib.RingTheory.Nakayama import Mathlib.RingTheory.Nilpotent.Basic import Mathlib.RingTheory.Nilpotent.Defs import Mathlib.RingTheory.Nilpotent.Lemmas +import Mathlib.RingTheory.Noetherian.Basic import Mathlib.RingTheory.Noetherian.Defs import Mathlib.RingTheory.Noetherian.Filter import Mathlib.RingTheory.Noetherian.Nilpotent diff --git a/Mathlib/Algebra/Module/FinitePresentation.lean b/Mathlib/Algebra/Module/FinitePresentation.lean index 553e1aee1ced2..b5737ac1b8e49 100644 --- a/Mathlib/Algebra/Module/FinitePresentation.lean +++ b/Mathlib/Algebra/Module/FinitePresentation.lean @@ -9,7 +9,7 @@ import Mathlib.LinearAlgebra.Isomorphisms import Mathlib.RingTheory.Finiteness.Projective import Mathlib.RingTheory.Finiteness.TensorProduct import Mathlib.RingTheory.Localization.BaseChange -import Mathlib.RingTheory.Noetherian.Defs +import Mathlib.RingTheory.Noetherian.Basic /-! diff --git a/Mathlib/FieldTheory/Tower.lean b/Mathlib/FieldTheory/Tower.lean index dbe2c6efe3f02..59cdb775f2ec4 100644 --- a/Mathlib/FieldTheory/Tower.lean +++ b/Mathlib/FieldTheory/Tower.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ -import Mathlib.RingTheory.Noetherian.Defs +import Mathlib.RingTheory.Noetherian.Basic /-! # Finiteness of `IsScalarTower` diff --git a/Mathlib/RingTheory/GradedAlgebra/Noetherian.lean b/Mathlib/RingTheory/GradedAlgebra/Noetherian.lean index cba8dab196628..9bd8cc5122295 100644 --- a/Mathlib/RingTheory/GradedAlgebra/Noetherian.lean +++ b/Mathlib/RingTheory/GradedAlgebra/Noetherian.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Fangming Li -/ import Mathlib.RingTheory.GradedAlgebra.Basic -import Mathlib.RingTheory.Noetherian.Defs +import Mathlib.RingTheory.Noetherian.Basic /-! # The properties of a graded Noetherian ring. diff --git a/Mathlib/RingTheory/Ideal/Quotient/Noetherian.lean b/Mathlib/RingTheory/Ideal/Quotient/Noetherian.lean index abe698078c199..2f2f268de1998 100644 --- a/Mathlib/RingTheory/Ideal/Quotient/Noetherian.lean +++ b/Mathlib/RingTheory/Ideal/Quotient/Noetherian.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ import Mathlib.RingTheory.Ideal.Quotient.Operations -import Mathlib.RingTheory.Noetherian.Defs +import Mathlib.RingTheory.Noetherian.Basic /-! # Noetherian quotient rings and quotient modules diff --git a/Mathlib/RingTheory/Noetherian/Basic.lean b/Mathlib/RingTheory/Noetherian/Basic.lean new file mode 100644 index 0000000000000..38156a143989d --- /dev/null +++ b/Mathlib/RingTheory/Noetherian/Basic.lean @@ -0,0 +1,366 @@ +/- +Copyright (c) 2018 Mario Carneiro, Kevin Buzzard. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro, Kevin Buzzard +-/ +import Mathlib.LinearAlgebra.Quotient.Basic +import Mathlib.RingTheory.Noetherian.Defs +import Mathlib.RingTheory.Finiteness.Cardinality +import Mathlib.RingTheory.Finiteness.Finsupp +import Mathlib.RingTheory.Ideal.Maps + +/-! +# Noetherian rings and modules + +The following are equivalent for a module M over a ring R: +1. Every increasing chain of submodules M₁ ⊆ M₂ ⊆ M₃ ⊆ ⋯ eventually stabilises. +2. Every submodule is finitely generated. + +A module satisfying these equivalent conditions is said to be a *Noetherian* R-module. +A ring is a *Noetherian ring* if it is Noetherian as a module over itself. + +(Note that we do not assume yet that our rings are commutative, +so perhaps this should be called "left Noetherian". +To avoid cumbersome names once we specialize to the commutative case, +we don't make this explicit in the declaration names.) + +## Main definitions + +Let `R` be a ring and let `M` and `P` be `R`-modules. Let `N` be an `R`-submodule of `M`. + +* `IsNoetherian R M` is the proposition that `M` is a Noetherian `R`-module. It is a class, + implemented as the predicate that all `R`-submodules of `M` are finitely generated. + +## Main statements + +* `isNoetherian_iff` is the theorem that an R-module M is Noetherian iff `>` is well-founded on + `Submodule R M`. + +Note that the Hilbert basis theorem, that if a commutative ring R is Noetherian then so is R[X], +is proved in `RingTheory.Polynomial`. + +## References + +* [M. F. Atiyah and I. G. Macdonald, *Introduction to commutative algebra*][atiyah-macdonald] +* [samuel1967] + +## Tags + +Noetherian, noetherian, Noetherian ring, Noetherian module, noetherian ring, noetherian module + +-/ + +assert_not_exists Matrix + +open Set Pointwise + +section + +variable {R : Type*} {M : Type*} {P : Type*} +variable [Semiring R] [AddCommMonoid M] [AddCommMonoid P] +variable [Module R M] [Module R P] + +open IsNoetherian + +variable (M) + +theorem isNoetherian_of_surjective (f : M →ₗ[R] P) (hf : LinearMap.range f = ⊤) [IsNoetherian R M] : + IsNoetherian R P := + ⟨fun s => + have : (s.comap f).map f = s := Submodule.map_comap_eq_self <| hf.symm ▸ le_top + this ▸ (noetherian _).map _⟩ + +variable {M} + +instance isNoetherian_range (f : M →ₗ[R] P) [IsNoetherian R M] : + IsNoetherian R (LinearMap.range f) := + isNoetherian_of_surjective _ _ f.range_rangeRestrict + +instance isNoetherian_quotient {A M : Type*} [Ring A] [AddCommGroup M] [SMul R A] [Module R M] + [Module A M] [IsScalarTower R A M] (N : Submodule A M) [IsNoetherian R M] : + IsNoetherian R (M ⧸ N) := + isNoetherian_of_surjective M ((Submodule.mkQ N).restrictScalars R) <| + LinearMap.range_eq_top.mpr N.mkQ_surjective + +@[deprecated (since := "2024-04-27"), nolint defLemma] +alias Submodule.Quotient.isNoetherian := isNoetherian_quotient + +theorem isNoetherian_of_linearEquiv (f : M ≃ₗ[R] P) [IsNoetherian R M] : IsNoetherian R P := + isNoetherian_of_surjective _ f.toLinearMap f.range + +theorem LinearEquiv.isNoetherian_iff (f : M ≃ₗ[R] P) : IsNoetherian R M ↔ IsNoetherian R P := + ⟨fun _ ↦ isNoetherian_of_linearEquiv f, fun _ ↦ isNoetherian_of_linearEquiv f.symm⟩ + +theorem isNoetherian_top_iff : IsNoetherian R (⊤ : Submodule R M) ↔ IsNoetherian R M := + Submodule.topEquiv.isNoetherian_iff + +theorem isNoetherian_of_injective [IsNoetherian R P] (f : M →ₗ[R] P) (hf : Function.Injective f) : + IsNoetherian R M := + isNoetherian_of_linearEquiv (LinearEquiv.ofInjective f hf).symm + +theorem fg_of_injective [IsNoetherian R P] {N : Submodule R M} (f : M →ₗ[R] P) + (hf : Function.Injective f) : N.FG := + haveI := isNoetherian_of_injective f hf + IsNoetherian.noetherian N + +end + +namespace Module + +variable {R M N : Type*} +variable [Semiring R] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] +variable (R M) + +-- see Note [lower instance priority] +instance (priority := 80) _root_.isNoetherian_of_finite [Finite M] : IsNoetherian R M := + ⟨fun s => ⟨(s : Set M).toFinite.toFinset, by rw [Set.Finite.coe_toFinset, Submodule.span_eq]⟩⟩ + +-- see Note [lower instance priority] +instance (priority := 100) IsNoetherian.finite [IsNoetherian R M] : Module.Finite R M := + ⟨IsNoetherian.noetherian ⊤⟩ + +instance {R₁ S : Type*} [CommSemiring R₁] [Semiring S] [Algebra R₁ S] + [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) : + Module.Finite R M := + ⟨fg_of_injective f hf⟩ + +end Module + +section + +variable {R : Type*} {M : Type*} {P : Type*} +variable [Ring R] [AddCommGroup M] [AddCommGroup P] +variable [Module R M] [Module R P] + +open IsNoetherian + +theorem isNoetherian_of_ker_bot [IsNoetherian R P] (f : M →ₗ[R] P) (hf : LinearMap.ker f = ⊥) : + IsNoetherian R M := + isNoetherian_of_linearEquiv (LinearEquiv.ofInjective f <| LinearMap.ker_eq_bot.mp hf).symm + +theorem fg_of_ker_bot [IsNoetherian R P] {N : Submodule R M} (f : M →ₗ[R] P) + (hf : LinearMap.ker f = ⊥) : N.FG := + haveI := isNoetherian_of_ker_bot f hf + IsNoetherian.noetherian N + +instance isNoetherian_prod [IsNoetherian R M] [IsNoetherian R P] : IsNoetherian R (M × P) := + ⟨fun s => + Submodule.fg_of_fg_map_of_fg_inf_ker (LinearMap.snd R M P) (noetherian _) <| + have : s ⊓ LinearMap.ker (LinearMap.snd R M P) ≤ LinearMap.range (LinearMap.inl R M P) := + fun x ⟨_, hx2⟩ => ⟨x.1, Prod.ext rfl <| Eq.symm <| LinearMap.mem_ker.1 hx2⟩ + Submodule.map_comap_eq_self this ▸ (noetherian _).map _⟩ + +instance isNoetherian_sup (M₁ M₂ : Submodule R P) [IsNoetherian R M₁] [IsNoetherian R M₂] : + IsNoetherian R ↥(M₁ ⊔ M₂) := by + have := isNoetherian_range (M₁.subtype.coprod M₂.subtype) + rwa [LinearMap.range_coprod, Submodule.range_subtype, Submodule.range_subtype] at this + +variable {ι : Type*} [Finite ι] + +instance isNoetherian_pi : + ∀ {M : ι → Type*} [∀ i, AddCommGroup (M i)] + [∀ i, Module R (M i)] [∀ i, IsNoetherian R (M i)], IsNoetherian R (∀ i, M i) := by + apply Finite.induction_empty_option _ _ _ ι + · exact fun e h ↦ isNoetherian_of_linearEquiv (LinearEquiv.piCongrLeft R _ e) + · infer_instance + · exact fun ih ↦ isNoetherian_of_linearEquiv (LinearEquiv.piOptionEquivProd R).symm + +/-- A version of `isNoetherian_pi` for non-dependent functions. We need this instance because +sometimes Lean fails to apply the dependent version in non-dependent settings (e.g., it fails to +prove that `ι → ℝ` is finite dimensional over `ℝ`). -/ +instance isNoetherian_pi' [IsNoetherian R M] : IsNoetherian R (ι → M) := + isNoetherian_pi + +instance isNoetherian_iSup : + ∀ {M : ι → Submodule R P} [∀ i, IsNoetherian R (M i)], IsNoetherian R ↥(⨆ i, M i) := by + apply Finite.induction_empty_option _ _ _ ι + · intro _ _ e h _ _; rw [← e.iSup_comp]; apply h + · intros; rw [iSup_of_empty]; infer_instance + · intro _ _ ih _ _; rw [iSup_option]; infer_instance + +end + +section CommRing + +variable (R M N : Type*) [CommRing R] [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] + [IsNoetherian R M] [Module.Finite R N] + +instance isNoetherian_linearMap_pi {ι : Type*} [Finite ι] : IsNoetherian R ((ι → R) →ₗ[R] M) := + let _i : Fintype ι := Fintype.ofFinite ι; isNoetherian_of_linearEquiv (Module.piEquiv ι R M) + +instance isNoetherian_linearMap : IsNoetherian R (N →ₗ[R] M) := by + obtain ⟨n, f, hf⟩ := Module.Finite.exists_fin' R N + let g : (N →ₗ[R] M) →ₗ[R] (Fin n → R) →ₗ[R] M := (LinearMap.llcomp R (Fin n → R) N M).flip f + exact isNoetherian_of_injective g hf.injective_linearMapComp_right + +end CommRing + +open IsNoetherian Submodule Function + +section + +universe w + +variable {R M P : Type*} {N : Type w} [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid N] + [Module R N] [AddCommMonoid P] [Module R P] + +/-- If `∀ I > J, P I` implies `P J`, then `P` holds for all submodules. -/ +theorem IsNoetherian.induction [IsNoetherian R M] {P : Submodule R M → Prop} + (hgt : ∀ I, (∀ J > I, P J) → P I) (I : Submodule R M) : P I := + IsWellFounded.induction _ I hgt + +end + +section + +universe w + +variable {R M P : Type*} {N : Type w} [Ring R] [AddCommGroup M] [Module R M] [AddCommGroup N] + [Module R N] [AddCommGroup P] [Module R P] [IsNoetherian R M] + +lemma Submodule.finite_ne_bot_of_iSupIndep {ι : Type*} {N : ι → Submodule R M} + (h : iSupIndep N) : + Set.Finite {i | N i ≠ ⊥} := + WellFoundedGT.finite_ne_bot_of_iSupIndep h + +@[deprecated (since := "2024-11-24")] +alias Submodule.finite_ne_bot_of_independent := Submodule.finite_ne_bot_of_iSupIndep + +/-- A linearly-independent family of vectors in a module over a non-trivial ring must be finite if +the module is Noetherian. -/ +theorem LinearIndependent.finite_of_isNoetherian [Nontrivial R] {ι} {v : ι → M} + (hv : LinearIndependent R v) : Finite ι := by + refine WellFoundedGT.finite_of_iSupIndep + hv.iSupIndep_span_singleton + fun i contra => ?_ + apply hv.ne_zero i + have : v i ∈ R ∙ v i := Submodule.mem_span_singleton_self (v i) + rwa [contra, Submodule.mem_bot] at this + +theorem LinearIndependent.set_finite_of_isNoetherian [Nontrivial R] {s : Set M} + (hi : LinearIndependent R ((↑) : s → M)) : s.Finite := + @Set.toFinite _ _ hi.finite_of_isNoetherian + +/-- If the first and final modules in an exact sequence are Noetherian, + then the middle module is also Noetherian. -/ +theorem isNoetherian_of_range_eq_ker [IsNoetherian R P] + (f : M →ₗ[R] N) (g : N →ₗ[R] P) (h : LinearMap.range f = LinearMap.ker g) : + IsNoetherian R N := + isNoetherian_mk <| + wellFounded_gt_exact_sequence + (LinearMap.range f) + (Submodule.map (f.ker.liftQ f le_rfl)) + (Submodule.comap (f.ker.liftQ f le_rfl)) + (Submodule.comap g.rangeRestrict) (Submodule.map g.rangeRestrict) + (Submodule.gciMapComap <| LinearMap.ker_eq_bot.mp <| Submodule.ker_liftQ_eq_bot _ _ _ le_rfl) + (Submodule.giMapComap g.surjective_rangeRestrict) + (by simp [Submodule.map_comap_eq, inf_comm, Submodule.range_liftQ]) + (by simp [Submodule.comap_map_eq, h]) + +theorem isNoetherian_iff_submodule_quotient (S : Submodule R P) : + IsNoetherian R P ↔ IsNoetherian R S ∧ IsNoetherian R (P ⧸ S) := by + refine ⟨fun _ ↦ ⟨inferInstance, inferInstance⟩, fun ⟨_, _⟩ ↦ ?_⟩ + apply isNoetherian_of_range_eq_ker S.subtype S.mkQ + rw [Submodule.ker_mkQ, Submodule.range_subtype] + +/-- A sequence `f` of submodules of a noetherian module, +with `f (n+1)` disjoint from the supremum of `f 0`, ..., `f n`, +is eventually zero. -/ +theorem IsNoetherian.disjoint_partialSups_eventually_bot + (f : ℕ → Submodule R M) (h : ∀ n, Disjoint (partialSups f n) (f (n + 1))) : + ∃ n : ℕ, ∀ m, n ≤ m → f m = ⊥ := by + -- A little off-by-one cleanup first: + suffices t : ∃ n : ℕ, ∀ m, n ≤ m → f (m + 1) = ⊥ by + obtain ⟨n, w⟩ := t + use n + 1 + rintro (_ | m) p + · cases p + · apply w + exact Nat.succ_le_succ_iff.mp p + obtain ⟨n, w⟩ := monotone_stabilizes_iff_noetherian.mpr inferInstance (partialSups f) + exact + ⟨n, fun m p => + (h m).eq_bot_of_ge <| sup_eq_left.1 <| (w (m + 1) <| le_add_right p).symm.trans <| w m p⟩ + +end + +-- see Note [lower instance priority] +/-- Modules over the trivial ring are Noetherian. -/ +instance (priority := 100) isNoetherian_of_subsingleton (R M) [Subsingleton R] [Semiring R] + [AddCommMonoid M] [Module R M] : IsNoetherian R M := + haveI := Module.subsingleton R M + isNoetherian_of_finite R M + +theorem isNoetherian_of_submodule_of_noetherian (R M) [Semiring R] [AddCommMonoid M] [Module R M] + (N : Submodule R M) (h : IsNoetherian R M) : IsNoetherian R N := + isNoetherian_mk ⟨OrderEmbedding.wellFounded (Submodule.MapSubtype.orderEmbedding N).dual h.wf⟩ + +/-- If `M / S / R` is a scalar tower, and `M / R` is Noetherian, then `M / S` is +also noetherian. -/ +theorem isNoetherian_of_tower (R) {S M} [Semiring R] [Semiring S] [AddCommMonoid M] [SMul R S] + [Module S M] [Module R M] [IsScalarTower R S M] (h : IsNoetherian R M) : IsNoetherian S M := + isNoetherian_mk ⟨(Submodule.restrictScalarsEmbedding R S M).dual.wellFounded h.wf⟩ + +theorem isNoetherian_of_fg_of_noetherian {R M} [Ring R] [AddCommGroup M] [Module R M] + (N : Submodule R M) [I : IsNoetherianRing R] (hN : N.FG) : IsNoetherian R N := by + let ⟨s, hs⟩ := hN + haveI := Classical.decEq M + haveI := Classical.decEq R + have : ∀ x ∈ s, x ∈ N := fun x hx => hs ▸ Submodule.subset_span hx + refine + @isNoetherian_of_surjective + R ((↑s : Set M) → R) N _ _ _ (Pi.module _ _ _) _ ?_ ?_ isNoetherian_pi + · fapply LinearMap.mk + · fapply AddHom.mk + · exact fun f => ⟨∑ i ∈ s.attach, f i • i.1, N.sum_mem fun c _ => N.smul_mem _ <| this _ c.2⟩ + · intro f g + apply Subtype.eq + change (∑ i ∈ s.attach, (f i + g i) • _) = _ + simp only [add_smul, Finset.sum_add_distrib] + rfl + · intro c f + apply Subtype.eq + change (∑ i ∈ s.attach, (c • f i) • _) = _ + simp only [smul_eq_mul, mul_smul] + exact Finset.smul_sum.symm + · rw [LinearMap.range_eq_top] + rintro ⟨n, hn⟩ + change n ∈ N at hn + rw [← hs, ← Set.image_id (s : Set M), Finsupp.mem_span_image_iff_linearCombination] at hn + rcases hn with ⟨l, hl1, hl2⟩ + refine ⟨fun x => l x, Subtype.ext ?_⟩ + change (∑ i ∈ s.attach, l i • (i : M)) = n + rw [s.sum_attach fun i ↦ l i • i, ← hl2, + Finsupp.linearCombination_apply, Finsupp.sum, eq_comm] + refine Finset.sum_subset hl1 fun x _ hx => ?_ + rw [Finsupp.not_mem_support_iff.1 hx, zero_smul] + +instance isNoetherian_of_isNoetherianRing_of_finite (R M : Type*) + [Ring R] [AddCommGroup M] [Module R M] [IsNoetherianRing R] [Module.Finite R M] : + IsNoetherian R M := + have : IsNoetherian R (⊤ : Submodule R M) := + isNoetherian_of_fg_of_noetherian _ <| Module.finite_def.mp inferInstance + isNoetherian_of_linearEquiv (LinearEquiv.ofTop (⊤ : Submodule R M) rfl) + +/-- In a module over a Noetherian ring, the submodule generated by finitely many vectors is +Noetherian. -/ +theorem isNoetherian_span_of_finite (R) {M} [Ring R] [AddCommGroup M] [Module R M] + [IsNoetherianRing R] {A : Set M} (hA : A.Finite) : IsNoetherian R (Submodule.span R A) := + isNoetherian_of_fg_of_noetherian _ (Submodule.fg_def.mpr ⟨A, hA, rfl⟩) + +theorem isNoetherianRing_of_surjective (R) [Ring R] (S) [Ring S] (f : R →+* S) + (hf : Function.Surjective f) [H : IsNoetherianRing R] : IsNoetherianRing S := + isNoetherian_mk ⟨OrderEmbedding.wellFounded (Ideal.orderEmbeddingOfSurjective f hf).dual H.wf⟩ + +instance isNoetherianRing_range {R} [Ring R] {S} [Ring S] (f : R →+* S) [IsNoetherianRing R] : + IsNoetherianRing f.range := + isNoetherianRing_of_surjective R f.range f.rangeRestrict f.rangeRestrict_surjective + +theorem isNoetherianRing_of_ringEquiv (R) [Ring R] {S} [Ring S] (f : R ≃+* S) [IsNoetherianRing R] : + IsNoetherianRing S := + isNoetherianRing_of_surjective R S f.toRingHom f.toEquiv.surjective diff --git a/Mathlib/RingTheory/Noetherian/Defs.lean b/Mathlib/RingTheory/Noetherian/Defs.lean index 926f9ba8c6880..e14e88932c6a3 100644 --- a/Mathlib/RingTheory/Noetherian/Defs.lean +++ b/Mathlib/RingTheory/Noetherian/Defs.lean @@ -3,10 +3,7 @@ Copyright (c) 2018 Mario Carneiro, Kevin Buzzard. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Kevin Buzzard -/ -import Mathlib.LinearAlgebra.Quotient.Basic -import Mathlib.RingTheory.Finiteness.Cardinality -import Mathlib.RingTheory.Finiteness.Finsupp -import Mathlib.RingTheory.Ideal.Maps +import Mathlib.RingTheory.Finiteness.Basic /-! # Noetherian rings and modules @@ -49,6 +46,9 @@ Noetherian, noetherian, Noetherian ring, Noetherian module, noetherian ring, noe -/ +assert_not_exists Finsupp.linearCombination +assert_not_exists Matrix +assert_not_exists Pi.basis open Set Pointwise @@ -100,144 +100,8 @@ theorem isNoetherian_of_le {s t : Submodule R M} [ht : IsNoetherian R t] (h : s IsNoetherian R s := isNoetherian_submodule.mpr fun _ hs' => isNoetherian_submodule.mp ht _ (le_trans hs' h) -variable (M) - -theorem isNoetherian_of_surjective (f : M →ₗ[R] P) (hf : LinearMap.range f = ⊤) [IsNoetherian R M] : - IsNoetherian R P := - ⟨fun s => - have : (s.comap f).map f = s := Submodule.map_comap_eq_self <| hf.symm ▸ le_top - this ▸ (noetherian _).map _⟩ - -variable {M} - -instance isNoetherian_range (f : M →ₗ[R] P) [IsNoetherian R M] : - IsNoetherian R (LinearMap.range f) := - isNoetherian_of_surjective _ _ f.range_rangeRestrict - -instance isNoetherian_quotient {A M : Type*} [Ring A] [AddCommGroup M] [SMul R A] [Module R M] - [Module A M] [IsScalarTower R A M] (N : Submodule A M) [IsNoetherian R M] : - IsNoetherian R (M ⧸ N) := - isNoetherian_of_surjective M ((Submodule.mkQ N).restrictScalars R) <| - LinearMap.range_eq_top.mpr N.mkQ_surjective - -@[deprecated (since := "2024-04-27"), nolint defLemma] -alias Submodule.Quotient.isNoetherian := isNoetherian_quotient - -theorem isNoetherian_of_linearEquiv (f : M ≃ₗ[R] P) [IsNoetherian R M] : IsNoetherian R P := - isNoetherian_of_surjective _ f.toLinearMap f.range - -theorem LinearEquiv.isNoetherian_iff (f : M ≃ₗ[R] P) : IsNoetherian R M ↔ IsNoetherian R P := - ⟨fun _ ↦ isNoetherian_of_linearEquiv f, fun _ ↦ isNoetherian_of_linearEquiv f.symm⟩ - -theorem isNoetherian_top_iff : IsNoetherian R (⊤ : Submodule R M) ↔ IsNoetherian R M := - Submodule.topEquiv.isNoetherian_iff - -theorem isNoetherian_of_injective [IsNoetherian R P] (f : M →ₗ[R] P) (hf : Function.Injective f) : - IsNoetherian R M := - isNoetherian_of_linearEquiv (LinearEquiv.ofInjective f hf).symm - -theorem fg_of_injective [IsNoetherian R P] {N : Submodule R M} (f : M →ₗ[R] P) - (hf : Function.Injective f) : N.FG := - haveI := isNoetherian_of_injective f hf - IsNoetherian.noetherian N - -end - -namespace Module - -variable {R M N : Type*} -variable [Semiring R] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] -variable (R M) - --- see Note [lower instance priority] -instance (priority := 80) _root_.isNoetherian_of_finite [Finite M] : IsNoetherian R M := - ⟨fun s => ⟨(s : Set M).toFinite.toFinset, by rw [Set.Finite.coe_toFinset, Submodule.span_eq]⟩⟩ - --- see Note [lower instance priority] -instance (priority := 100) IsNoetherian.finite [IsNoetherian R M] : Module.Finite R M := - ⟨IsNoetherian.noetherian ⊤⟩ - -instance {R₁ S : Type*} [CommSemiring R₁] [Semiring S] [Algebra R₁ S] - [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) : - Module.Finite R M := - ⟨fg_of_injective f hf⟩ - -end Module - -section - -variable {R : Type*} {M : Type*} {P : Type*} -variable [Ring R] [AddCommGroup M] [AddCommGroup P] -variable [Module R M] [Module R P] - -open IsNoetherian - -theorem isNoetherian_of_ker_bot [IsNoetherian R P] (f : M →ₗ[R] P) (hf : LinearMap.ker f = ⊥) : - IsNoetherian R M := - isNoetherian_of_linearEquiv (LinearEquiv.ofInjective f <| LinearMap.ker_eq_bot.mp hf).symm - -theorem fg_of_ker_bot [IsNoetherian R P] {N : Submodule R M} (f : M →ₗ[R] P) - (hf : LinearMap.ker f = ⊥) : N.FG := - haveI := isNoetherian_of_ker_bot f hf - IsNoetherian.noetherian N - -instance isNoetherian_prod [IsNoetherian R M] [IsNoetherian R P] : IsNoetherian R (M × P) := - ⟨fun s => - Submodule.fg_of_fg_map_of_fg_inf_ker (LinearMap.snd R M P) (noetherian _) <| - have : s ⊓ LinearMap.ker (LinearMap.snd R M P) ≤ LinearMap.range (LinearMap.inl R M P) := - fun x ⟨_, hx2⟩ => ⟨x.1, Prod.ext rfl <| Eq.symm <| LinearMap.mem_ker.1 hx2⟩ - Submodule.map_comap_eq_self this ▸ (noetherian _).map _⟩ - -instance isNoetherian_sup (M₁ M₂ : Submodule R P) [IsNoetherian R M₁] [IsNoetherian R M₂] : - IsNoetherian R ↥(M₁ ⊔ M₂) := by - have := isNoetherian_range (M₁.subtype.coprod M₂.subtype) - rwa [LinearMap.range_coprod, Submodule.range_subtype, Submodule.range_subtype] at this - -variable {ι : Type*} [Finite ι] - -instance isNoetherian_pi : - ∀ {M : ι → Type*} [∀ i, AddCommGroup (M i)] - [∀ i, Module R (M i)] [∀ i, IsNoetherian R (M i)], IsNoetherian R (∀ i, M i) := by - apply Finite.induction_empty_option _ _ _ ι - · exact fun e h ↦ isNoetherian_of_linearEquiv (LinearEquiv.piCongrLeft R _ e) - · infer_instance - · exact fun ih ↦ isNoetherian_of_linearEquiv (LinearEquiv.piOptionEquivProd R).symm - -/-- A version of `isNoetherian_pi` for non-dependent functions. We need this instance because -sometimes Lean fails to apply the dependent version in non-dependent settings (e.g., it fails to -prove that `ι → ℝ` is finite dimensional over `ℝ`). -/ -instance isNoetherian_pi' [IsNoetherian R M] : IsNoetherian R (ι → M) := - isNoetherian_pi - -instance isNoetherian_iSup : - ∀ {M : ι → Submodule R P} [∀ i, IsNoetherian R (M i)], IsNoetherian R ↥(⨆ i, M i) := by - apply Finite.induction_empty_option _ _ _ ι - · intro _ _ e h _ _; rw [← e.iSup_comp]; apply h - · intros; rw [iSup_of_empty]; infer_instance - · intro _ _ ih _ _; rw [iSup_option]; infer_instance - end -section CommRing - -variable (R M N : Type*) [CommRing R] [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] - [IsNoetherian R M] [Module.Finite R N] - -instance isNoetherian_linearMap_pi {ι : Type*} [Finite ι] : IsNoetherian R ((ι → R) →ₗ[R] M) := - let _i : Fintype ι := Fintype.ofFinite ι; isNoetherian_of_linearEquiv (Module.piEquiv ι R M) - -instance isNoetherian_linearMap : IsNoetherian R (N →ₗ[R] M) := by - obtain ⟨n, f, hf⟩ := Module.Finite.exists_fin' R N - let g : (N →ₗ[R] M) →ₗ[R] (Fin n → R) →ₗ[R] M := (LinearMap.llcomp R (Fin n → R) N M).flip f - exact isNoetherian_of_injective g hf.injective_linearMapComp_right - -end CommRing - open IsNoetherian Submodule Function section @@ -299,84 +163,6 @@ theorem monotone_stabilizes_iff_noetherian : (∀ f : ℕ →o Submodule R M, ∃ n, ∀ m, n ≤ m → f n = f m) ↔ IsNoetherian R M := by rw [isNoetherian_iff, WellFounded.monotone_chain_condition] -/-- If `∀ I > J, P I` implies `P J`, then `P` holds for all submodules. -/ -theorem IsNoetherian.induction [IsNoetherian R M] {P : Submodule R M → Prop} - (hgt : ∀ I, (∀ J > I, P J) → P I) (I : Submodule R M) : P I := - IsWellFounded.induction _ I hgt - -end - -section - -universe w - -variable {R M P : Type*} {N : Type w} [Ring R] [AddCommGroup M] [Module R M] [AddCommGroup N] - [Module R N] [AddCommGroup P] [Module R P] [IsNoetherian R M] - -lemma Submodule.finite_ne_bot_of_iSupIndep {ι : Type*} {N : ι → Submodule R M} - (h : iSupIndep N) : - Set.Finite {i | N i ≠ ⊥} := - WellFoundedGT.finite_ne_bot_of_iSupIndep h - -@[deprecated (since := "2024-11-24")] -alias Submodule.finite_ne_bot_of_independent := Submodule.finite_ne_bot_of_iSupIndep - -/-- A linearly-independent family of vectors in a module over a non-trivial ring must be finite if -the module is Noetherian. -/ -theorem LinearIndependent.finite_of_isNoetherian [Nontrivial R] {ι} {v : ι → M} - (hv : LinearIndependent R v) : Finite ι := by - refine WellFoundedGT.finite_of_iSupIndep - hv.iSupIndep_span_singleton - fun i contra => ?_ - apply hv.ne_zero i - have : v i ∈ R ∙ v i := Submodule.mem_span_singleton_self (v i) - rwa [contra, Submodule.mem_bot] at this - -theorem LinearIndependent.set_finite_of_isNoetherian [Nontrivial R] {s : Set M} - (hi : LinearIndependent R ((↑) : s → M)) : s.Finite := - @Set.toFinite _ _ hi.finite_of_isNoetherian - -/-- If the first and final modules in an exact sequence are Noetherian, - then the middle module is also Noetherian. -/ -theorem isNoetherian_of_range_eq_ker [IsNoetherian R P] - (f : M →ₗ[R] N) (g : N →ₗ[R] P) (h : LinearMap.range f = LinearMap.ker g) : - IsNoetherian R N := - isNoetherian_mk <| - wellFounded_gt_exact_sequence - (LinearMap.range f) - (Submodule.map (f.ker.liftQ f le_rfl)) - (Submodule.comap (f.ker.liftQ f le_rfl)) - (Submodule.comap g.rangeRestrict) (Submodule.map g.rangeRestrict) - (Submodule.gciMapComap <| LinearMap.ker_eq_bot.mp <| Submodule.ker_liftQ_eq_bot _ _ _ le_rfl) - (Submodule.giMapComap g.surjective_rangeRestrict) - (by simp [Submodule.map_comap_eq, inf_comm, Submodule.range_liftQ]) - (by simp [Submodule.comap_map_eq, h]) - -theorem isNoetherian_iff_submodule_quotient (S : Submodule R P) : - IsNoetherian R P ↔ IsNoetherian R S ∧ IsNoetherian R (P ⧸ S) := by - refine ⟨fun _ ↦ ⟨inferInstance, inferInstance⟩, fun ⟨_, _⟩ ↦ ?_⟩ - apply isNoetherian_of_range_eq_ker S.subtype S.mkQ - rw [Submodule.ker_mkQ, Submodule.range_subtype] - -/-- A sequence `f` of submodules of a noetherian module, -with `f (n+1)` disjoint from the supremum of `f 0`, ..., `f n`, -is eventually zero. -/ -theorem IsNoetherian.disjoint_partialSups_eventually_bot - (f : ℕ → Submodule R M) (h : ∀ n, Disjoint (partialSups f n) (f (n + 1))) : - ∃ n : ℕ, ∀ m, n ≤ m → f m = ⊥ := by - -- A little off-by-one cleanup first: - suffices t : ∃ n : ℕ, ∀ m, n ≤ m → f (m + 1) = ⊥ by - obtain ⟨n, w⟩ := t - use n + 1 - rintro (_ | m) p - · cases p - · apply w - exact Nat.succ_le_succ_iff.mp p - obtain ⟨n, w⟩ := monotone_stabilizes_iff_noetherian.mpr inferInstance (partialSups f) - exact - ⟨n, fun m p => - (h m).eq_bot_of_ge <| sup_eq_left.1 <| (w (m + 1) <| le_add_right p).symm.trans <| w m p⟩ - end /-- A (semi)ring is Noetherian if it is Noetherian as a module over itself, @@ -391,79 +177,3 @@ theorem isNoetherianRing_iff {R} [Semiring R] : IsNoetherianRing R ↔ IsNoether theorem isNoetherianRing_iff_ideal_fg (R : Type*) [Semiring R] : IsNoetherianRing R ↔ ∀ I : Ideal R, I.FG := isNoetherianRing_iff.trans isNoetherian_def - --- see Note [lower instance priority] -/-- Modules over the trivial ring are Noetherian. -/ -instance (priority := 100) isNoetherian_of_subsingleton (R M) [Subsingleton R] [Semiring R] - [AddCommMonoid M] [Module R M] : IsNoetherian R M := - haveI := Module.subsingleton R M - isNoetherian_of_finite R M - -theorem isNoetherian_of_submodule_of_noetherian (R M) [Semiring R] [AddCommMonoid M] [Module R M] - (N : Submodule R M) (h : IsNoetherian R M) : IsNoetherian R N := - isNoetherian_mk ⟨OrderEmbedding.wellFounded (Submodule.MapSubtype.orderEmbedding N).dual h.wf⟩ - -/-- If `M / S / R` is a scalar tower, and `M / R` is Noetherian, then `M / S` is -also noetherian. -/ -theorem isNoetherian_of_tower (R) {S M} [Semiring R] [Semiring S] [AddCommMonoid M] [SMul R S] - [Module S M] [Module R M] [IsScalarTower R S M] (h : IsNoetherian R M) : IsNoetherian S M := - isNoetherian_mk ⟨(Submodule.restrictScalarsEmbedding R S M).dual.wellFounded h.wf⟩ - -theorem isNoetherian_of_fg_of_noetherian {R M} [Ring R] [AddCommGroup M] [Module R M] - (N : Submodule R M) [I : IsNoetherianRing R] (hN : N.FG) : IsNoetherian R N := by - let ⟨s, hs⟩ := hN - haveI := Classical.decEq M - haveI := Classical.decEq R - have : ∀ x ∈ s, x ∈ N := fun x hx => hs ▸ Submodule.subset_span hx - refine - @isNoetherian_of_surjective - R ((↑s : Set M) → R) N _ _ _ (Pi.module _ _ _) _ ?_ ?_ isNoetherian_pi - · fapply LinearMap.mk - · fapply AddHom.mk - · exact fun f => ⟨∑ i ∈ s.attach, f i • i.1, N.sum_mem fun c _ => N.smul_mem _ <| this _ c.2⟩ - · intro f g - apply Subtype.eq - change (∑ i ∈ s.attach, (f i + g i) • _) = _ - simp only [add_smul, Finset.sum_add_distrib] - rfl - · intro c f - apply Subtype.eq - change (∑ i ∈ s.attach, (c • f i) • _) = _ - simp only [smul_eq_mul, mul_smul] - exact Finset.smul_sum.symm - · rw [LinearMap.range_eq_top] - rintro ⟨n, hn⟩ - change n ∈ N at hn - rw [← hs, ← Set.image_id (s : Set M), Finsupp.mem_span_image_iff_linearCombination] at hn - rcases hn with ⟨l, hl1, hl2⟩ - refine ⟨fun x => l x, Subtype.ext ?_⟩ - change (∑ i ∈ s.attach, l i • (i : M)) = n - rw [s.sum_attach fun i ↦ l i • i, ← hl2, - Finsupp.linearCombination_apply, Finsupp.sum, eq_comm] - refine Finset.sum_subset hl1 fun x _ hx => ?_ - rw [Finsupp.not_mem_support_iff.1 hx, zero_smul] - -instance isNoetherian_of_isNoetherianRing_of_finite (R M : Type*) - [Ring R] [AddCommGroup M] [Module R M] [IsNoetherianRing R] [Module.Finite R M] : - IsNoetherian R M := - have : IsNoetherian R (⊤ : Submodule R M) := - isNoetherian_of_fg_of_noetherian _ <| Module.finite_def.mp inferInstance - isNoetherian_of_linearEquiv (LinearEquiv.ofTop (⊤ : Submodule R M) rfl) - -/-- In a module over a Noetherian ring, the submodule generated by finitely many vectors is -Noetherian. -/ -theorem isNoetherian_span_of_finite (R) {M} [Ring R] [AddCommGroup M] [Module R M] - [IsNoetherianRing R] {A : Set M} (hA : A.Finite) : IsNoetherian R (Submodule.span R A) := - isNoetherian_of_fg_of_noetherian _ (Submodule.fg_def.mpr ⟨A, hA, rfl⟩) - -theorem isNoetherianRing_of_surjective (R) [Ring R] (S) [Ring S] (f : R →+* S) - (hf : Function.Surjective f) [H : IsNoetherianRing R] : IsNoetherianRing S := - isNoetherian_mk ⟨OrderEmbedding.wellFounded (Ideal.orderEmbeddingOfSurjective f hf).dual H.wf⟩ - -instance isNoetherianRing_range {R} [Ring R] {S} [Ring S] (f : R →+* S) [IsNoetherianRing R] : - IsNoetherianRing f.range := - isNoetherianRing_of_surjective R f.range f.rangeRestrict f.rangeRestrict_surjective - -theorem isNoetherianRing_of_ringEquiv (R) [Ring R] {S} [Ring S] (f : R ≃+* S) [IsNoetherianRing R] : - IsNoetherianRing S := - isNoetherianRing_of_surjective R S f.toRingHom f.toEquiv.surjective diff --git a/Mathlib/RingTheory/Noetherian/Orzech.lean b/Mathlib/RingTheory/Noetherian/Orzech.lean index 8858b55202f80..da0b31cf04f08 100644 --- a/Mathlib/RingTheory/Noetherian/Orzech.lean +++ b/Mathlib/RingTheory/Noetherian/Orzech.lean @@ -6,7 +6,7 @@ Authors: Mario Carneiro, Kevin Buzzard import Mathlib.Algebra.Module.Submodule.IterateMapComap import Mathlib.Algebra.Order.Archimedean.Basic import Mathlib.Order.PartialSups -import Mathlib.RingTheory.Noetherian.Defs +import Mathlib.RingTheory.Noetherian.Basic import Mathlib.RingTheory.OrzechProperty import Mathlib.Order.Filter.AtTopBot diff --git a/Mathlib/RingTheory/Polynomial/Basic.lean b/Mathlib/RingTheory/Polynomial/Basic.lean index 99c10175f0d49..4bee44c7de8c5 100644 --- a/Mathlib/RingTheory/Polynomial/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Basic.lean @@ -8,7 +8,7 @@ import Mathlib.Algebra.GeomSum import Mathlib.Algebra.MvPolynomial.CommRing import Mathlib.Algebra.MvPolynomial.Equiv import Mathlib.Algebra.Polynomial.BigOperators -import Mathlib.RingTheory.Noetherian.Defs +import Mathlib.RingTheory.Noetherian.Basic /-! # Ring-theoretic supplement of Algebra.Polynomial. diff --git a/Mathlib/RingTheory/PrimeSpectrum.lean b/Mathlib/RingTheory/PrimeSpectrum.lean index 29c25b19c2cf9..dd7f6aa053423 100644 --- a/Mathlib/RingTheory/PrimeSpectrum.lean +++ b/Mathlib/RingTheory/PrimeSpectrum.lean @@ -7,7 +7,7 @@ import Mathlib.LinearAlgebra.Finsupp.SumProd import Mathlib.RingTheory.Ideal.Prod import Mathlib.RingTheory.Localization.Ideal import Mathlib.RingTheory.Nilpotent.Lemmas -import Mathlib.RingTheory.Noetherian.Defs +import Mathlib.RingTheory.Noetherian.Basic /-! # Prime spectrum of a commutative (semi)ring as a type diff --git a/Mathlib/RingTheory/PrincipalIdealDomain.lean b/Mathlib/RingTheory/PrincipalIdealDomain.lean index 3a675457d9104..5343a17c2253a 100644 --- a/Mathlib/RingTheory/PrincipalIdealDomain.lean +++ b/Mathlib/RingTheory/PrincipalIdealDomain.lean @@ -5,6 +5,7 @@ Authors: Chris Hughes, Morenikeji Neri -/ import Mathlib.Algebra.EuclideanDomain.Field import Mathlib.Algebra.GCDMonoid.Basic +import Mathlib.RingTheory.Ideal.Maps import Mathlib.RingTheory.Ideal.Nonunits import Mathlib.RingTheory.Noetherian.UniqueFactorizationDomain From 3fa5f452ee1652a47dcf496bc3d52db2f2bdd810 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 25 Nov 2024 13:34:38 +0000 Subject: [PATCH 288/829] chore: ensure there is a replacement identifier or text suggestion for all deprecations (#19426) This is for compatibility with the stricter requirements of leanprover/lean4#6112. Anyone is welcome to replace the "No deprecation message was provided." messages in other PRs! :-) --- Mathlib/Algebra/BigOperators/Group/List.lean | 2 +- Mathlib/Algebra/Group/Action/Defs.lean | 8 +- Mathlib/Algebra/Group/AddChar.lean | 2 +- Mathlib/Algebra/Group/Equiv/Basic.lean | 20 +++- Mathlib/Algebra/Group/Pointwise/Set/Card.lean | 29 +++++- Mathlib/Algebra/Module/LinearMap/Defs.lean | 4 +- Mathlib/Algebra/Order/Field/Power.lean | 4 +- .../Algebra/Order/Group/Pointwise/Bounds.lean | 7 +- Mathlib/Analysis/Complex/Circle.lean | 6 +- Mathlib/CategoryTheory/Adjunction/Limits.lean | 8 +- .../CategoryTheory/Idempotents/Karoubi.lean | 2 +- Mathlib/CategoryTheory/Limits/Creates.lean | 12 +-- .../Limits/FunctorCategory/Basic.lean | 8 +- Mathlib/Data/Complex/Abs.lean | 2 +- Mathlib/Data/Fin/Basic.lean | 2 +- Mathlib/Data/List/Basic.lean | 8 +- Mathlib/Data/List/Flatten.lean | 3 +- Mathlib/Data/List/Indexes.lean | 12 +-- Mathlib/Data/List/Lex.lean | 2 +- Mathlib/Data/List/Pairwise.lean | 3 +- Mathlib/Data/List/Permutation.lean | 2 +- Mathlib/Data/Nat/Defs.lean | 2 +- Mathlib/Data/Set/Pointwise/SMul.lean | 9 +- Mathlib/Data/Setoid/Basic.lean | 4 +- Mathlib/Deprecated/AlgebraClasses.lean | 47 ++++----- Mathlib/Deprecated/ByteArray.lean | 10 +- Mathlib/Deprecated/Combinator.lean | 9 +- Mathlib/Deprecated/Equiv.lean | 4 +- Mathlib/Deprecated/LazyList.lean | 10 +- Mathlib/Deprecated/Logic.lean | 74 +++++++-------- Mathlib/Deprecated/NatLemmas.lean | 8 +- Mathlib/Deprecated/RelClasses.lean | 10 +- Mathlib/GroupTheory/Coset/Basic.lean | 8 +- Mathlib/GroupTheory/Coset/Defs.lean | 20 +++- Mathlib/GroupTheory/Exponent.lean | 3 +- Mathlib/GroupTheory/GroupAction/Defs.lean | 8 +- Mathlib/GroupTheory/GroupAction/Quotient.lean | 17 +++- Mathlib/GroupTheory/OrderOfElement.lean | 7 +- .../GroupTheory/SpecificGroups/Cyclic.lean | 9 +- Mathlib/LinearAlgebra/BilinearForm/Basic.lean | 10 +- Mathlib/LinearAlgebra/BilinearForm/Hom.lean | 20 ++-- Mathlib/LinearAlgebra/Prod.lean | 28 +++--- Mathlib/Logic/Basic.lean | 9 +- Mathlib/MeasureTheory/Function/L1Space.lean | 2 +- .../MeasureTheory/Measure/Typeclasses.lean | 2 +- .../Measure/WithDensityFinite.lean | 2 +- Mathlib/NumberTheory/ArithmeticFunction.lean | 3 +- Mathlib/NumberTheory/MulChar/Basic.lean | 6 +- Mathlib/Order/Defs.lean | 8 +- Mathlib/Order/OmegaCompletePartialOrder.lean | 2 +- Mathlib/Order/RelClasses.lean | 6 +- Mathlib/SetTheory/Cardinal/Aleph.lean | 18 ++-- Mathlib/SetTheory/Cardinal/Basic.lean | 2 +- Mathlib/SetTheory/Cardinal/Cofinality.lean | 2 +- Mathlib/SetTheory/Game/Ordinal.lean | 2 +- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 95 ++++++++++--------- Mathlib/SetTheory/Ordinal/Basic.lean | 8 +- Mathlib/SetTheory/Ordinal/Enum.lean | 2 +- Mathlib/SetTheory/Ordinal/FixedPoint.lean | 66 ++++++------- Mathlib/SetTheory/Ordinal/Principal.lean | 4 +- Mathlib/SetTheory/ZFC/Basic.lean | 43 +++++---- Mathlib/Topology/Algebra/ConstMulAction.lean | 2 +- .../Topology/Algebra/ContinuousMonoidHom.lean | 3 +- Mathlib/Topology/Algebra/Group/Quotient.lean | 6 +- Mathlib/Topology/LocalAtTarget.lean | 4 +- Mathlib/Topology/MetricSpace/Polish.lean | 2 +- Mathlib/Topology/NoetherianSpace.lean | 2 +- Mathlib/Topology/Order/OrderClosed.lean | 2 +- 68 files changed, 442 insertions(+), 324 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index bf01489af4f70..2bd6c8a351d66 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -634,7 +634,7 @@ end MonoidHom end MonoidHom set_option linter.deprecated false in -@[simp, deprecated (since := "2024-10-17")] +@[simp, deprecated "No deprecation message was provided." (since := "2024-10-17")] lemma Nat.sum_eq_listSum (l : List ℕ) : Nat.sum l = l.sum := rfl namespace List diff --git a/Mathlib/Algebra/Group/Action/Defs.lean b/Mathlib/Algebra/Group/Action/Defs.lean index 823dbf69f3586..294281be04be5 100644 --- a/Mathlib/Algebra/Group/Action/Defs.lean +++ b/Mathlib/Algebra/Group/Action/Defs.lean @@ -297,9 +297,15 @@ lemma smul_mul_smul_comm [Mul α] [Mul β] [SMul α β] [IsScalarTower α β β] (a • b) * (c • d) = (a * c) • (b * d) := by have : SMulCommClass β α β := .symm ..; exact smul_smul_smul_comm a b c d -@[to_additive (attr := deprecated (since := "2024-08-29"))] +@[to_additive] alias smul_mul_smul := smul_mul_smul_comm +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated smul_mul_smul_comm (since := "2024-08-29")] smul_mul_smul +attribute [deprecated vadd_add_vadd_comm (since := "2024-08-29")] vadd_add_vadd + + /-- Note that the `IsScalarTower α β β` and `SMulCommClass α β β` typeclass arguments are usually satisfied by `Algebra α β`. -/ @[to_additive] diff --git a/Mathlib/Algebra/Group/AddChar.lean b/Mathlib/Algebra/Group/AddChar.lean index 4b0c0f3e8d9bd..0279bc188d39a 100644 --- a/Mathlib/Algebra/Group/AddChar.lean +++ b/Mathlib/Algebra/Group/AddChar.lean @@ -253,7 +253,7 @@ lemma ne_one_iff : ψ ≠ 1 ↔ ∃ x, ψ x ≠ 1 := DFunLike.ne_iff lemma ne_zero_iff : ψ ≠ 0 ↔ ∃ x, ψ x ≠ 1 := DFunLike.ne_iff /-- An additive character is *nontrivial* if it takes a value `≠ 1`. -/ -@[deprecated (since := "2024-06-06")] +@[deprecated "No deprecation message was provided." (since := "2024-06-06")] def IsNontrivial (ψ : AddChar A M) : Prop := ∃ a : A, ψ a ≠ 1 set_option linter.deprecated false in diff --git a/Mathlib/Algebra/Group/Equiv/Basic.lean b/Mathlib/Algebra/Group/Equiv/Basic.lean index 203ef93e84ead..b0b66218245be 100644 --- a/Mathlib/Algebra/Group/Equiv/Basic.lean +++ b/Mathlib/Algebra/Group/Equiv/Basic.lean @@ -94,12 +94,28 @@ class MulEquivClass (F : Type*) (A B : outParam Type*) [Mul A] [Mul B] [EquivLik /-- Preserves multiplication. -/ map_mul : ∀ (f : F) (a b), f (a * b) = f a * f b -@[to_additive (attr := deprecated (since := "2024-11-10"))] +@[to_additive] alias MulEquivClass.map_eq_one_iff := EmbeddingLike.map_eq_one_iff -@[to_additive (attr := deprecated (since := "2024-11-10"))] +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated EmbeddingLike.map_eq_one_iff (since := "2024-11-10")] +MulEquivClass.map_eq_one_iff +attribute [deprecated EmbeddingLike.map_eq_zero_iff (since := "2024-11-10")] +AddEquivClass.map_eq_zero_iff + + +@[to_additive] alias MulEquivClass.map_ne_one_iff := EmbeddingLike.map_ne_one_iff +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated EmbeddingLike.map_ne_one_iff (since := "2024-11-10")] +MulEquivClass.map_ne_one_iff +attribute [deprecated EmbeddingLike.map_ne_zero_iff (since := "2024-11-10")] +AddEquivClass.map_ne_zero_iff + + namespace MulEquivClass variable (F) diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean index 404fdd46e25ad..d31f9773753a2 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean @@ -33,7 +33,13 @@ lemma natCard_mul_le : Nat.card (s * t) ≤ Nat.card s * Nat.card t := by refine Cardinal.toNat_le_toNat Cardinal.mk_mul_le ?_ aesop (add simp [Cardinal.mul_lt_aleph0_iff, finite_mul]) -@[to_additive (attr := deprecated (since := "2024-09-30"))] alias card_mul_le := natCard_mul_le +@[to_additive] alias card_mul_le := natCard_mul_le + +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated natCard_mul_le (since := "2024-09-30")] card_mul_le +attribute [deprecated natCard_add_le (since := "2024-09-30")] card_add_le + end Mul @@ -48,7 +54,12 @@ lemma _root_.Cardinal.mk_inv (s : Set G) : #↥(s⁻¹) = #s := by lemma natCard_inv (s : Set G) : Nat.card ↥(s⁻¹) = Nat.card s := by rw [← image_inv, Nat.card_image_of_injective inv_injective] -@[to_additive (attr := deprecated (since := "2024-09-30"))] alias card_inv := natCard_inv +@[to_additive] alias card_inv := natCard_inv + +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated natCard_inv (since := "2024-09-30")] card_inv +attribute [deprecated natCard_neg (since := "2024-09-30")] card_neg @[to_additive (attr := simp)] lemma encard_inv (s : Set G) : s⁻¹.encard = s.encard := by simp [encard, PartENat.card] @@ -74,7 +85,12 @@ variable [Group G] {s t : Set G} lemma natCard_div_le : Nat.card (s / t) ≤ Nat.card s * Nat.card t := by rw [div_eq_mul_inv, ← natCard_inv t]; exact natCard_mul_le -@[to_additive (attr := deprecated (since := "2024-09-30"))] alias card_div_le := natCard_div_le +@[to_additive] alias card_div_le := natCard_div_le + +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated natCard_div_le (since := "2024-09-30")] card_div_le +attribute [deprecated natCard_sub_le (since := "2024-09-30")] card_sub_le variable [MulAction G α] @@ -86,9 +102,14 @@ lemma _root_.Cardinal.mk_smul_set (a : G) (s : Set α) : #↥(a • s) = #s := lemma natCard_smul_set (a : G) (s : Set α) : Nat.card ↥(a • s) = Nat.card s := Nat.card_image_of_injective (MulAction.injective a) _ -@[to_additive (attr := deprecated (since := "2024-09-30"))] +@[to_additive] alias card_smul_set := Cardinal.mk_smul_set +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated Cardinal.mk_smul_set (since := "2024-09-30")] card_smul_set +attribute [deprecated Cardinal.mk_vadd_set (since := "2024-09-30")] card_vadd_set + @[to_additive (attr := simp)] lemma encard_smul_set (a : G) (s : Set α) : (a • s).encard = s.encard := by simp [encard, PartENat.card] diff --git a/Mathlib/Algebra/Module/LinearMap/Defs.lean b/Mathlib/Algebra/Module/LinearMap/Defs.lean index c6436036fb7c1..d4c25f58c52fb 100644 --- a/Mathlib/Algebra/Module/LinearMap/Defs.lean +++ b/Mathlib/Algebra/Module/LinearMap/Defs.lean @@ -596,14 +596,14 @@ variable [Semiring R] [Module R M] [Semiring S] [Module S M₂] [Module R M₃] variable {σ : R →+* S} /-- A `DistribMulActionHom` between two modules is a linear map. -/ -@[deprecated (since := "2024-11-08")] +@[deprecated "No deprecation message was provided." (since := "2024-11-08")] def toSemilinearMap (fₗ : M →ₑ+[σ.toMonoidHom] M₂) : M →ₛₗ[σ] M₂ := { fₗ with } instance : SemilinearMapClass (M →ₑ+[σ.toMonoidHom] M₂) σ M M₂ where /-- A `DistribMulActionHom` between two modules is a linear map. -/ -@[deprecated (since := "2024-11-08")] +@[deprecated "No deprecation message was provided." (since := "2024-11-08")] def toLinearMap (fₗ : M →+[R] M₃) : M →ₗ[R] M₃ := { fₗ with } diff --git a/Mathlib/Algebra/Order/Field/Power.lean b/Mathlib/Algebra/Order/Field/Power.lean index d363f2170a5c4..9cec9ad4fcc29 100644 --- a/Mathlib/Algebra/Order/Field/Power.lean +++ b/Mathlib/Algebra/Order/Field/Power.lean @@ -73,13 +73,13 @@ theorem zpow_injective (h₀ : 0 < a) (h₁ : a ≠ 1) : Injective (a ^ · : ℤ theorem zpow_inj (h₀ : 0 < a) (h₁ : a ≠ 1) : a ^ m = a ^ n ↔ m = n := zpow_right_inj₀ h₀ h₁ -@[deprecated (since := "2024-10-08")] +@[deprecated "No deprecation message was provided." (since := "2024-10-08")] theorem zpow_le_max_of_min_le {x : α} (hx : 1 ≤ x) {a b c : ℤ} (h : min a b ≤ c) : x ^ (-c) ≤ max (x ^ (-a)) (x ^ (-b)) := have : Antitone fun n : ℤ => x ^ (-n) := fun _ _ h => zpow_le_zpow_right₀ hx (neg_le_neg h) (this h).trans_eq this.map_min -@[deprecated (since := "2024-10-08")] +@[deprecated "No deprecation message was provided." (since := "2024-10-08")] theorem zpow_le_max_iff_min_le {x : α} (hx : 1 < x) {a b c : ℤ} : x ^ (-c) ≤ max (x ^ (-a)) (x ^ (-b)) ↔ min a b ≤ c := by simp_rw [le_max_iff, min_le_iff, zpow_le_zpow_iff_right₀ hx, neg_le_neg_iff] diff --git a/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean b/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean index ceb329876b8e6..294f6be7c92df 100644 --- a/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean +++ b/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean @@ -48,7 +48,12 @@ lemma BddAbove.mul (hs : BddAbove s) (ht : BddAbove t) : BddAbove (s * t) := lemma BddBelow.mul (hs : BddBelow s) (ht : BddBelow t) : BddBelow (s * t) := (Nonempty.mul hs ht).mono (subset_lowerBounds_mul s t) -@[to_additive (attr := deprecated (since := "2024-11-13"))] alias Set.BddAbove.mul := BddAbove.mul +@[to_additive] alias Set.BddAbove.mul := BddAbove.mul + +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated BddAbove.mul (since := "2024-11-13")] Set.BddAbove.mul +attribute [deprecated BddAbove.add (since := "2024-11-13")] Set.BddAbove.add @[to_additive] lemma BddAbove.range_mul (hf : BddAbove (range f)) (hg : BddAbove (range g)) : diff --git a/Mathlib/Analysis/Complex/Circle.lean b/Mathlib/Analysis/Complex/Circle.lean index 72e66ae46c1b7..e4a9a2e283278 100644 --- a/Mathlib/Analysis/Complex/Circle.lean +++ b/Mathlib/Analysis/Complex/Circle.lean @@ -41,17 +41,17 @@ open ComplexConjugate /-- The unit circle in `ℂ`, here given the structure of a submonoid of `ℂ`. Please use `Circle` when referring to the circle as a type. -/ -@[deprecated (since := "2024-07-24")] +@[deprecated "No deprecation message was provided." (since := "2024-07-24")] def circle : Submonoid ℂ := Submonoid.unitSphere ℂ set_option linter.deprecated false in -@[deprecated (since := "2024-07-24")] +@[deprecated "No deprecation message was provided." (since := "2024-07-24")] theorem mem_circle_iff_abs {z : ℂ} : z ∈ circle ↔ abs z = 1 := mem_sphere_zero_iff_norm set_option linter.deprecated false in -@[deprecated (since := "2024-07-24")] +@[deprecated "No deprecation message was provided." (since := "2024-07-24")] theorem mem_circle_iff_normSq {z : ℂ} : z ∈ circle ↔ normSq z = 1 := by simp [Complex.abs, mem_circle_iff_abs] diff --git a/Mathlib/CategoryTheory/Adjunction/Limits.lean b/Mathlib/CategoryTheory/Adjunction/Limits.lean index dd86591976534..2db2c795ef2c2 100644 --- a/Mathlib/CategoryTheory/Adjunction/Limits.lean +++ b/Mathlib/CategoryTheory/Adjunction/Limits.lean @@ -92,7 +92,7 @@ lemma leftAdjoint_preservesColimits : PreservesColimitsOfSize.{v, u} F where ((adj.functorialityAdjunction _).homEquiv _ _)⟩ } } include adj in -@[deprecated (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (since := "2024-11-19")] lemma leftAdjointPreservesColimits : PreservesColimitsOfSize.{v, u} F := adj.leftAdjoint_preservesColimits @@ -117,7 +117,7 @@ noncomputable instance (priority := 100) { reflects := fun t => ⟨(isColimitOfPreserves E.inv t).mapCoconeEquiv E.asEquivalence.unitIso.symm⟩ } } -@[deprecated (since := "2024-11-18")] +@[deprecated "No deprecation message was provided." (since := "2024-11-18")] lemma isEquivalenceReflectsColimits (E : D ⥤ C) [E.IsEquivalence] : ReflectsColimitsOfSize.{v, u} E := Functor.reflectsColimits_of_isEquivalence E @@ -216,7 +216,7 @@ lemma rightAdjoint_preservesLimits : PreservesLimitsOfSize.{v, u} G where ((adj.functorialityAdjunction' _).homEquiv _ _).symm⟩ } } include adj in -@[deprecated (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (since := "2024-11-19")] lemma rightAdjointPreservesLimits : PreservesLimitsOfSize.{v, u} G := adj.rightAdjoint_preservesLimits @@ -240,7 +240,7 @@ noncomputable instance (priority := 100) { reflects := fun t => ⟨(isLimitOfPreserves E.inv t).mapConeEquiv E.asEquivalence.unitIso.symm⟩ } } -@[deprecated (since := "2024-11-18")] +@[deprecated "No deprecation message was provided." (since := "2024-11-18")] lemma isEquivalenceReflectsLimits (E : D ⥤ C) [E.IsEquivalence] : ReflectsLimitsOfSize.{v, u} E := Functor.reflectsLimits_of_isEquivalence E diff --git a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean index 1c0f0936e00c6..d4b4fc8a4f69e 100644 --- a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean +++ b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean @@ -111,7 +111,7 @@ theorem comp_f {P Q R : Karoubi C} (f : P ⟶ Q) (g : Q ⟶ R) : (f ≫ g).f = f @[simp] theorem id_f {P : Karoubi C} : Hom.f (𝟙 P) = P.p := rfl -@[deprecated (since := "2024-07-15")] +@[deprecated "No deprecation message was provided." (since := "2024-07-15")] theorem id_eq {P : Karoubi C} : 𝟙 P = ⟨P.p, by repeat' rw [P.idem]⟩ := rfl /-- It is possible to coerce an object of `C` into an object of `Karoubi C`. diff --git a/Mathlib/CategoryTheory/Limits/Creates.lean b/Mathlib/CategoryTheory/Limits/Creates.lean index 148eafbe8efa1..9769ea2520000 100644 --- a/Mathlib/CategoryTheory/Limits/Creates.lean +++ b/Mathlib/CategoryTheory/Limits/Creates.lean @@ -332,7 +332,7 @@ instance (priority := 100) preservesLimit_of_createsLimit_and_hasLimit (K : J ((liftedLimitMapsToOriginal (limit.isLimit _)).symm ≪≫ (Cones.functoriality K F).mapIso ((liftedLimitIsLimit (limit.isLimit _)).uniqueUpToIso t))⟩ -@[deprecated (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (since := "2024-11-19")] lemma preservesLimitOfCreatesLimitAndHasLimit (K : J ⥤ C) (F : C ⥤ D) [CreatesLimit K F] [HasLimit (K ⋙ F)] : PreservesLimit K F := preservesLimit_of_createsLimit_and_hasLimit _ _ @@ -342,7 +342,7 @@ lemma preservesLimitOfCreatesLimitAndHasLimit (K : J ⥤ C) (F : C ⥤ D) instance (priority := 100) preservesLimitOfShape_of_createsLimitsOfShape_and_hasLimitsOfShape (F : C ⥤ D) [CreatesLimitsOfShape J F] [HasLimitsOfShape J D] : PreservesLimitsOfShape J F where -@[deprecated (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (since := "2024-11-19")] lemma preservesLimitOfShapeOfCreatesLimitsOfShapeAndHasLimitsOfShape (F : C ⥤ D) [CreatesLimitsOfShape J F] [HasLimitsOfShape J D] : PreservesLimitsOfShape J F := @@ -354,7 +354,7 @@ instance (priority := 100) preservesLimits_of_createsLimits_and_hasLimits (F : C [CreatesLimitsOfSize.{w, w'} F] [HasLimitsOfSize.{w, w'} D] : PreservesLimitsOfSize.{w, w'} F where -@[deprecated (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (since := "2024-11-19")] lemma preservesLimitsOfCreatesLimitsAndHasLimits (F : C ⥤ D) [CreatesLimitsOfSize.{w, w'} F] [HasLimitsOfSize.{w, w'} D] : PreservesLimitsOfSize.{w, w'} F := @@ -466,7 +466,7 @@ instance (priority := 100) preservesColimit_of_createsColimit_and_hasColimit (K (Cocones.functoriality K F).mapIso ((liftedColimitIsColimit (colimit.isColimit _)).uniqueUpToIso t))⟩ -@[deprecated (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (since := "2024-11-19")] lemma preservesColimitOfCreatesColimitAndHasColimit (K : J ⥤ C) (F : C ⥤ D) [CreatesColimit K F] [HasColimit (K ⋙ F)] : PreservesColimit K F := preservesColimit_of_createsColimit_and_hasColimit _ _ @@ -477,7 +477,7 @@ instance (priority := 100) preservesColimitOfShape_of_createsColimitsOfShape_and (F : C ⥤ D) [CreatesColimitsOfShape J F] [HasColimitsOfShape J D] : PreservesColimitsOfShape J F where -@[deprecated (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (since := "2024-11-19")] lemma preservesColimitOfShapeOfCreatesColimitsOfShapeAndHasColimitsOfShape (F : C ⥤ D) [CreatesColimitsOfShape J F] [HasColimitsOfShape J D] : PreservesColimitsOfShape J F := @@ -489,7 +489,7 @@ instance (priority := 100) preservesColimits_of_createsColimits_and_hasColimits [CreatesColimitsOfSize.{w, w'} F] [HasColimitsOfSize.{w, w'} D] : PreservesColimitsOfSize.{w, w'} F where -@[deprecated (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (since := "2024-11-19")] lemma preservesColimitsOfCreatesColimitsAndHasColimits (F : C ⥤ D) [CreatesColimitsOfSize.{w, w'} F] [HasColimitsOfSize.{w, w'} D] : PreservesColimitsOfSize.{w, w'} F := diff --git a/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean b/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean index 357916b1fe01e..603001e9e6ce6 100644 --- a/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean +++ b/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean @@ -362,7 +362,7 @@ lemma preservesLimit_of_evaluation (F : D ⥤ K ⥤ C) (G : J ⥤ D) change IsLimit ((F ⋙ (evaluation K C).obj X).mapCone c) exact isLimitOfPreserves _ hc⟩⟩ -@[deprecated (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (since := "2024-11-19")] lemma preservesLimitOfEvaluation (F : D ⥤ K ⥤ C) (G : J ⥤ D) (H : ∀ k : K, PreservesLimit G (F ⋙ (evaluation K C).obj k : D ⥤ C)) : PreservesLimit G F := @@ -374,7 +374,7 @@ lemma preservesLimitsOfShape_of_evaluation (F : D ⥤ K ⥤ C) (J : Type*) [Cate PreservesLimitsOfShape J F := ⟨fun {G} => preservesLimit_of_evaluation F G fun _ => PreservesLimitsOfShape.preservesLimit⟩ -@[deprecated (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (since := "2024-11-19")] lemma preservesLimitsOfShapeOfEvaluation (F : D ⥤ K ⥤ C) (J : Type*) [Category J] (H : ∀ k : K, PreservesLimitsOfShape J (F ⋙ (evaluation K C).obj k)) : PreservesLimitsOfShape J F := @@ -387,7 +387,7 @@ lemma preservesLimits_of_evaluation (F : D ⥤ K ⥤ C) ⟨fun {L} _ => preservesLimitsOfShape_of_evaluation F L fun _ => PreservesLimitsOfSize.preservesLimitsOfShape⟩ -@[deprecated (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (since := "2024-11-19")] lemma preservesLimitsOfEvaluation (F : D ⥤ K ⥤ C) (H : ∀ k : K, PreservesLimitsOfSize.{w', w} (F ⋙ (evaluation K C).obj k)) : PreservesLimitsOfSize.{w', w} F := @@ -412,7 +412,7 @@ lemma preservesColimit_of_evaluation (F : D ⥤ K ⥤ C) (G : J ⥤ D) change IsColimit ((F ⋙ (evaluation K C).obj X).mapCocone c) exact isColimitOfPreserves _ hc⟩⟩ -@[deprecated (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (since := "2024-11-19")] lemma preservesColimitOfEvaluation (F : D ⥤ K ⥤ C) (G : J ⥤ D) (H : ∀ k, PreservesColimit G (F ⋙ (evaluation K C).obj k)) : PreservesColimit G F := preservesColimit_of_evaluation _ _ H diff --git a/Mathlib/Data/Complex/Abs.lean b/Mathlib/Data/Complex/Abs.lean index 99319c9d298b3..18f2f48e1b525 100644 --- a/Mathlib/Data/Complex/Abs.lean +++ b/Mathlib/Data/Complex/Abs.lean @@ -199,7 +199,7 @@ theorem abs_im_div_abs_le_one (z : ℂ) : |z.im / Complex.abs z| ≤ 1 := @[simp, norm_cast] lemma abs_intCast (n : ℤ) : abs n = |↑n| := by rw [← ofReal_intCast, abs_ofReal] -@[deprecated (since := "2024-02-14")] +@[deprecated "No deprecation message was provided." (since := "2024-02-14")] lemma int_cast_abs (n : ℤ) : |↑n| = Complex.abs n := (abs_intCast _).symm theorem normSq_eq_abs (x : ℂ) : normSq x = (Complex.abs x) ^ 2 := by diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 21c9e27353985..4fcd646907769 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -1115,7 +1115,7 @@ lemma succAbove_ne_last {a : Fin (n + 2)} {b : Fin (n + 1)} (ha : a ≠ last _) lemma succAbove_last_apply (i : Fin n) : succAbove (last n) i = castSucc i := by rw [succAbove_last] -@[deprecated (since := "2024-05-30")] +@[deprecated "No deprecation message was provided." (since := "2024-05-30")] lemma succAbove_lt_ge (p : Fin (n + 1)) (i : Fin n) : castSucc i < p ∨ p ≤ castSucc i := Nat.lt_or_ge (castSucc i) p diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 4e3ce8533a890..3f51c9f665d12 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -34,7 +34,7 @@ universe u v w variable {ι : Type*} {α : Type u} {β : Type v} {γ : Type w} {l₁ l₂ : List α} /-- `≤` implies not `>` for lists. -/ -@[deprecated (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] theorem le_eq_not_gt [LT α] : ∀ l₁ l₂ : List α, (l₁ ≤ l₂) = ¬l₂ < l₁ := fun _ _ => rfl -- Porting note: Delete this attribute @@ -496,7 +496,7 @@ theorem get_tail (l : List α) (i) (h : i < l.tail.length) l.tail.get ⟨i, h⟩ = l.get ⟨i + 1, h'⟩ := by cases l <;> [cases h; rfl] -@[deprecated (since := "2024-08-22")] +@[deprecated "No deprecation message was provided." (since := "2024-08-22")] theorem get_cons {l : List α} {a : α} {n} (hl) : (a :: l).get ⟨n, hl⟩ = if hn : n = 0 then a else l.get ⟨n - 1, by contrapose! hl; rw [length_cons]; omega⟩ := @@ -627,7 +627,7 @@ lemma cons_sublist_cons' {a b : α} : a :: l₁ <+ b :: l₂ ↔ a :: l₁ <+ l theorem sublist_cons_of_sublist (a : α) (h : l₁ <+ l₂) : l₁ <+ a :: l₂ := h.cons _ -@[deprecated (since := "2024-04-07")] +@[deprecated "No deprecation message was provided." (since := "2024-04-07")] theorem sublist_of_cons_sublist_cons {a} (h : a :: l₁ <+ a :: l₂) : l₁ <+ l₂ := h.of_cons_cons @[deprecated (since := "2024-04-07")] alias cons_sublist_cons_iff := cons_sublist_cons @@ -2281,7 +2281,7 @@ theorem length_dropSlice_lt (i j : ℕ) (hj : 0 < j) (xs : List α) (hi : i < xs simp; omega set_option linter.deprecated false in -@[deprecated (since := "2024-07-25")] +@[deprecated "No deprecation message was provided." (since := "2024-07-25")] theorem sizeOf_dropSlice_lt [SizeOf α] (i j : ℕ) (hj : 0 < j) (xs : List α) (hi : i < xs.length) : SizeOf.sizeOf (List.dropSlice i j xs) < SizeOf.sizeOf xs := by induction xs generalizing i j hj with diff --git a/Mathlib/Data/List/Flatten.lean b/Mathlib/Data/List/Flatten.lean index f6fbc8e3044fc..797ba9dc5c297 100644 --- a/Mathlib/Data/List/Flatten.lean +++ b/Mathlib/Data/List/Flatten.lean @@ -158,7 +158,8 @@ theorem append_flatten_map_append (L : List (List α)) (x : List α) : @[deprecated (since := "2024-10-15")] alias append_join_map_append := append_flatten_map_append -@[deprecated (since := "2024-08-15")] theorem sublist_join {l} {L : List (List α)} (h : l ∈ L) : +@[deprecated "No deprecation message was provided." (since := "2024-08-15")] +theorem sublist_join {l} {L : List (List α)} (h : l ∈ L) : l <+ L.flatten := sublist_flatten_of_mem h diff --git a/Mathlib/Data/List/Indexes.lean b/Mathlib/Data/List/Indexes.lean index a1b935a65e0cc..fdbe9168bfa47 100644 --- a/Mathlib/Data/List/Indexes.lean +++ b/Mathlib/Data/List/Indexes.lean @@ -89,7 +89,7 @@ theorem mapIdx_eq_ofFn (l : List α) (f : ℕ → α → β) : section deprecated /-- Lean3 `map_with_index` helper function -/ -@[deprecated (since := "2024-08-15")] +@[deprecated "No deprecation message was provided." (since := "2024-08-15")] protected def oldMapIdxCore (f : ℕ → α → β) : ℕ → List α → List β | _, [] => [] | k, a :: as => f k a :: List.oldMapIdxCore f (k + 1) as @@ -97,12 +97,12 @@ protected def oldMapIdxCore (f : ℕ → α → β) : ℕ → List α → List set_option linter.deprecated false in /-- Given a function `f : ℕ → α → β` and `as : List α`, `as = [a₀, a₁, ...]`, returns the list `[f 0 a₀, f 1 a₁, ...]`. -/ -@[deprecated (since := "2024-08-15")] +@[deprecated "No deprecation message was provided." (since := "2024-08-15")] protected def oldMapIdx (f : ℕ → α → β) (as : List α) : List β := List.oldMapIdxCore f 0 as set_option linter.deprecated false in -@[deprecated (since := "2024-08-15")] +@[deprecated "No deprecation message was provided." (since := "2024-08-15")] protected theorem oldMapIdxCore_eq (l : List α) (f : ℕ → α → β) (n : ℕ) : l.oldMapIdxCore f n = l.oldMapIdx fun i a ↦ f (i + n) a := by induction' l with hd tl hl generalizing f n @@ -111,7 +111,7 @@ protected theorem oldMapIdxCore_eq (l : List α) (f : ℕ → α → β) (n : simp only [List.oldMapIdxCore, hl, Nat.add_left_comm, Nat.add_comm, Nat.add_zero] set_option linter.deprecated false in -@[deprecated (since := "2024-08-15")] +@[deprecated "No deprecation message was provided." (since := "2024-08-15")] protected theorem oldMapIdxCore_append : ∀ (f : ℕ → α → β) (n : ℕ) (l₁ l₂ : List α), List.oldMapIdxCore f n (l₁ ++ l₂) = List.oldMapIdxCore f n l₁ ++ List.oldMapIdxCore f (n + l₁.length) l₂ := by @@ -139,7 +139,7 @@ protected theorem oldMapIdxCore_append : ∀ (f : ℕ → α → β) (n : ℕ) ( rw [Nat.add_assoc]; simp only [Nat.add_comm] set_option linter.deprecated false in -@[deprecated (since := "2024-08-15")] +@[deprecated "No deprecation message was provided." (since := "2024-08-15")] protected theorem oldMapIdx_append : ∀ (f : ℕ → α → β) (l : List α) (e : α), List.oldMapIdx f (l ++ [e]) = List.oldMapIdx f l ++ [f l.length e] := by intros f l e @@ -148,7 +148,7 @@ protected theorem oldMapIdx_append : ∀ (f : ℕ → α → β) (l : List α) ( simp only [Nat.zero_add]; rfl set_option linter.deprecated false in -@[deprecated (since := "2024-08-15")] +@[deprecated "No deprecation message was provided." (since := "2024-08-15")] protected theorem new_def_eq_old_def : ∀ (f : ℕ → α → β) (l : List α), l.mapIdx f = List.oldMapIdx f l := by intro f diff --git a/Mathlib/Data/List/Lex.lean b/Mathlib/Data/List/Lex.lean index fef097d96a69f..69566d9aa5d91 100644 --- a/Mathlib/Data/List/Lex.lean +++ b/Mathlib/Data/List/Lex.lean @@ -99,7 +99,7 @@ instance isAsymm (r : α → α → Prop) [IsAsymm α r] : IsAsymm (List α) (Le | _, _, Lex.cons _, Lex.rel h₂ => asymm h₂ h₂ | _, _, Lex.cons h₁, Lex.cons h₂ => aux _ _ h₁ h₂ -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] instance isStrictTotalOrder (r : α → α → Prop) [IsStrictTotalOrder α r] : IsStrictTotalOrder (List α) (Lex r) := { isStrictWeakOrder_of_isOrderConnected with } diff --git a/Mathlib/Data/List/Pairwise.lean b/Mathlib/Data/List/Pairwise.lean index 4682916a610a0..be4f87bb84be6 100644 --- a/Mathlib/Data/List/Pairwise.lean +++ b/Mathlib/Data/List/Pairwise.lean @@ -50,7 +50,8 @@ theorem Pairwise.set_pairwise (hl : Pairwise R l) (hr : Symmetric R) : { x | x hl.forall hr -- Porting note: Duplicate of `pairwise_map` but with `f` explicit. -@[deprecated (since := "2024-02-25")] theorem pairwise_map' (f : β → α) : +@[deprecated "No deprecation message was provided." (since := "2024-02-25")] +theorem pairwise_map' (f : β → α) : ∀ {l : List β}, Pairwise R (map f l) ↔ Pairwise (R on f) l | [] => by simp only [map, Pairwise.nil] | b :: l => by diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index 2ce7b4ce525f3..3e0073e3e2049 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -419,7 +419,7 @@ theorem length_permutations'Aux (s : List α) (x : α) : · simp · simpa using IH -@[deprecated (since := "2024-06-12")] +@[deprecated "No deprecation message was provided." (since := "2024-06-12")] theorem permutations'Aux_get_zero (s : List α) (x : α) (hn : 0 < length (permutations'Aux x s) := (by simp)) : (permutations'Aux x s).get ⟨0, hn⟩ = x :: s := diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index 85ee5070c28b1..20c3db731c333 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -1092,7 +1092,7 @@ lemma sub_mod_eq_zero_of_mod_eq (h : m % k = n % k) : (m - n) % k = 0 := by lemma one_mod_eq_one : ∀ {n : ℕ}, 1 % n = 1 ↔ n ≠ 1 | 0 | 1 | n + 2 => by simp -@[deprecated (since := "2024-08-28")] +@[deprecated "No deprecation message was provided." (since := "2024-08-28")] lemma one_mod_of_ne_one : ∀ {n : ℕ}, n ≠ 1 → 1 % n = 1 := one_mod_eq_one.mpr lemma dvd_sub_mod (k : ℕ) : n ∣ k - k % n := diff --git a/Mathlib/Data/Set/Pointwise/SMul.lean b/Mathlib/Data/Set/Pointwise/SMul.lean index 1f4a5da3be5cd..e7008d4c67fbf 100644 --- a/Mathlib/Data/Set/Pointwise/SMul.lean +++ b/Mathlib/Data/Set/Pointwise/SMul.lean @@ -415,8 +415,13 @@ lemma disjoint_smul_set_left : Disjoint (a • s) t ↔ Disjoint s (a⁻¹ • t lemma disjoint_smul_set_right : Disjoint s (a • t) ↔ Disjoint (a⁻¹ • s) t := by simpa using disjoint_smul_set (a := a) (s := a⁻¹ • s) -@[to_additive (attr := deprecated (since := "2024-10-18"))] -alias smul_set_disjoint_iff := disjoint_smul_set +@[to_additive] alias smul_set_disjoint_iff := disjoint_smul_set + +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated disjoint_smul_set (since := "2024-10-18")] smul_set_disjoint_iff +attribute [deprecated disjoint_vadd_set (since := "2024-10-18")] vadd_set_disjoint_iff + end Group diff --git a/Mathlib/Data/Setoid/Basic.lean b/Mathlib/Data/Setoid/Basic.lean index 2e75e96025a44..bbaa5b239fbfe 100644 --- a/Mathlib/Data/Setoid/Basic.lean +++ b/Mathlib/Data/Setoid/Basic.lean @@ -36,12 +36,12 @@ attribute [trans] Setoid.trans variable {α : Type*} {β : Type*} /-- A version of `Setoid.r` that takes the equivalence relation as an explicit argument. -/ -@[deprecated (since := "2024-08-29")] +@[deprecated "No deprecation message was provided." (since := "2024-08-29")] def Setoid.Rel (r : Setoid α) : α → α → Prop := @Setoid.r _ r set_option linter.deprecated false in -@[deprecated (since := "2024-10-09")] +@[deprecated "No deprecation message was provided." (since := "2024-10-09")] instance Setoid.decidableRel (r : Setoid α) [h : DecidableRel r.r] : DecidableRel r.Rel := h diff --git a/Mathlib/Deprecated/AlgebraClasses.lean b/Mathlib/Deprecated/AlgebraClasses.lean index cc26f497422b2..bb6f948871913 100644 --- a/Mathlib/Deprecated/AlgebraClasses.lean +++ b/Mathlib/Deprecated/AlgebraClasses.lean @@ -25,16 +25,16 @@ universe u v variable {α : Sort u} -@[deprecated (since := "2024-09-11")] +@[deprecated "No deprecation message was provided." (since := "2024-09-11")] class IsLeftCancel (α : Sort u) (op : α → α → α) : Prop where left_cancel : ∀ a b c, op a b = op a c → b = c -@[deprecated (since := "2024-09-11")] +@[deprecated "No deprecation message was provided." (since := "2024-09-11")] class IsRightCancel (α : Sort u) (op : α → α → α) : Prop where right_cancel : ∀ a b c, op a b = op c b → a = c /-- `IsTotalPreorder X r` means that the binary relation `r` on `X` is total and a preorder. -/ -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] class IsTotalPreorder (α : Sort u) (r : α → α → Prop) extends IsTrans α r, IsTotal α r : Prop /-- Every total pre-order is a pre-order. -/ @@ -45,11 +45,11 @@ instance (priority := 100) isTotalPreorder_isPreorder (α : Sort u) (r : α → /-- `IsIncompTrans X lt` means that for `lt` a binary relation on `X`, the incomparable relation `fun a b => ¬ lt a b ∧ ¬ lt b a` is transitive. -/ -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] class IsIncompTrans (α : Sort u) (lt : α → α → Prop) : Prop where incomp_trans : ∀ a b c, ¬lt a b ∧ ¬lt b a → ¬lt b c ∧ ¬lt c b → ¬lt a c ∧ ¬lt c a -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] instance (priority := 100) (α : Sort u) (lt : α → α → Prop) [IsStrictWeakOrder α lt] : IsIncompTrans α lt := { ‹IsStrictWeakOrder α lt› with } @@ -59,7 +59,7 @@ variable {r : α → α → Prop} local infixl:50 " ≺ " => r -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem incomp_trans [IsIncompTrans α r] {a b c : α} : ¬a ≺ b ∧ ¬b ≺ a → ¬b ≺ c ∧ ¬c ≺ b → ¬a ≺ c ∧ ¬c ≺ a := IsIncompTrans.incomp_trans _ _ _ @@ -68,7 +68,8 @@ section ExplicitRelationVariants variable (r) -@[elab_without_expected_type, deprecated (since := "2024-07-30")] +@[elab_without_expected_type, + deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem incomp_trans_of [IsIncompTrans α r] {a b c : α} : ¬a ≺ b ∧ ¬b ≺ a → ¬b ≺ c ∧ ¬c ≺ b → ¬a ≺ c ∧ ¬c ≺ a := incomp_trans @@ -85,32 +86,32 @@ variable {r : α → α → Prop} local infixl:50 " ≺ " => r -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] def Equiv (a b : α) : Prop := ¬a ≺ b ∧ ¬b ≺ a local infixl:50 " ≈ " => @Equiv _ r -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem esymm {a b : α} : a ≈ b → b ≈ a := fun ⟨h₁, h₂⟩ => ⟨h₂, h₁⟩ -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem not_lt_of_equiv {a b : α} : a ≈ b → ¬a ≺ b := fun h => h.1 -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem not_lt_of_equiv' {a b : α} : a ≈ b → ¬b ≺ a := fun h => h.2 variable [IsStrictWeakOrder α r] -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem erefl (a : α) : a ≈ a := ⟨irrefl a, irrefl a⟩ -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem etrans {a b c : α} : a ≈ b → b ≈ c → a ≈ c := incomp_trans -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] instance isEquiv : IsEquiv α (@Equiv _ r) where refl := erefl trans _ _ _ := etrans @@ -123,7 +124,7 @@ notation:50 a " ≈[" lt "]" b:50 => @Equiv _ lt a b end StrictWeakOrder -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem isStrictWeakOrder_of_isTotalPreorder {α : Sort u} {le : α → α → Prop} {lt : α → α → Prop} [DecidableRel le] [IsTotalPreorder α le] (h : ∀ a b, lt a b ↔ ¬le b a) : IsStrictWeakOrder α lt := @@ -149,20 +150,20 @@ section LinearOrder variable {α : Type*} [LinearOrder α] set_option linter.deprecated false in -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] instance : IsTotalPreorder α (· ≤ ·) where trans := @le_trans _ _ total := le_total set_option linter.deprecated false in -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] instance isStrictWeakOrder_of_linearOrder : IsStrictWeakOrder α (· < ·) := have : IsTotalPreorder α (· ≤ ·) := by infer_instance -- Porting note: added isStrictWeakOrder_of_isTotalPreorder lt_iff_not_ge end LinearOrder -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem lt_of_lt_of_incomp {α : Sort u} {lt : α → α → Prop} [IsStrictWeakOrder α lt] [DecidableRel lt] : ∀ {a b c}, lt a b → ¬lt b c ∧ ¬lt c b → lt a c := @fun a b c hab ⟨nbc, ncb⟩ => @@ -171,7 +172,7 @@ theorem lt_of_lt_of_incomp {α : Sort u} {lt : α → α → Prop} [IsStrictWeak have : ¬lt a b ∧ ¬lt b a := incomp_trans_of lt ⟨nac, nca⟩ ⟨ncb, nbc⟩ absurd hab this.1 -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem lt_of_incomp_of_lt {α : Sort u} {lt : α → α → Prop} [IsStrictWeakOrder α lt] [DecidableRel lt] : ∀ {a b c}, ¬lt a b ∧ ¬lt b a → lt b c → lt a c := @fun a b c ⟨nab, nba⟩ hbc => @@ -180,7 +181,7 @@ theorem lt_of_incomp_of_lt {α : Sort u} {lt : α → α → Prop} [IsStrictWeak have : ¬lt b c ∧ ¬lt c b := incomp_trans_of lt ⟨nba, nab⟩ ⟨nac, nca⟩ absurd hbc this.1 -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem eq_of_incomp {α : Sort u} {lt : α → α → Prop} [IsTrichotomous α lt] {a b} : ¬lt a b ∧ ¬lt b a → a = b := fun ⟨nab, nba⟩ => match trichotomous_of lt a b with @@ -188,17 +189,17 @@ theorem eq_of_incomp {α : Sort u} {lt : α → α → Prop} [IsTrichotomous α | Or.inr (Or.inl hab) => hab | Or.inr (Or.inr hba) => absurd hba nba -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem eq_of_eqv_lt {α : Sort u} {lt : α → α → Prop} [IsTrichotomous α lt] {a b} : a ≈[lt]b → a = b := eq_of_incomp -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem incomp_iff_eq {α : Sort u} {lt : α → α → Prop} [IsTrichotomous α lt] [IsIrrefl α lt] (a b) : ¬lt a b ∧ ¬lt b a ↔ a = b := Iff.intro eq_of_incomp fun hab => hab ▸ And.intro (irrefl_of lt a) (irrefl_of lt a) -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem eqv_lt_iff_eq {α : Sort u} {lt : α → α → Prop} [IsTrichotomous α lt] [IsIrrefl α lt] (a b) : a ≈[lt]b ↔ a = b := incomp_iff_eq a b diff --git a/Mathlib/Deprecated/ByteArray.lean b/Mathlib/Deprecated/ByteArray.lean index c199be41b3f7c..b7146dbf289c6 100644 --- a/Mathlib/Deprecated/ByteArray.lean +++ b/Mathlib/Deprecated/ByteArray.lean @@ -19,7 +19,7 @@ set_option linter.deprecated false namespace Nat /-- A well-ordered relation for "upwards" induction on the natural numbers up to some bound `ub`. -/ -@[deprecated (since := "2024-08-19")] +@[deprecated "No deprecation message was provided." (since := "2024-08-19")] def Up (ub a i : Nat) := i < a ∧ i < ub theorem Up.next {ub i} (h : i < ub) : Up ub (i+1) i := ⟨Nat.lt_succ_self _, h⟩ @@ -28,13 +28,13 @@ theorem Up.WF (ub) : WellFounded (Up ub) := Subrelation.wf (h₂ := (measure (ub - ·)).wf) fun ⟨ia, iu⟩ ↦ Nat.sub_lt_sub_left iu ia /-- A well-ordered relation for "upwards" induction on the natural numbers up to some bound `ub`. -/ -@[deprecated (since := "2024-08-19")] +@[deprecated "No deprecation message was provided." (since := "2024-08-19")] def upRel (ub : Nat) : WellFoundedRelation Nat := ⟨Up ub, Up.WF ub⟩ end Nat /-- A terminal byte slice, a suffix of a byte array. -/ -@[deprecated (since := "2024-08-19")] +@[deprecated "No deprecation message was provided." (since := "2024-08-19")] structure ByteSliceT := (arr : ByteArray) (off : Nat) namespace ByteSliceT @@ -66,7 +66,7 @@ def toArray : ByteSlice → ByteArray universe u v /-- The inner loop of the `forIn` implementation for byte slices. -/ -@[deprecated (since := "2024-08-19")] +@[deprecated "No deprecation message was provided." (since := "2024-08-19")] def forIn.loop {m : Type u → Type v} {β : Type u} [Monad m] (f : UInt8 → β → m (ForInStep β)) (arr : ByteArray) (off _end : Nat) (i : Nat) (b : β) : m β := if h : i < _end then do @@ -75,7 +75,7 @@ def forIn.loop {m : Type u → Type v} {β : Type u} [Monad m] (f : UInt8 → β | ForInStep.yield b => have := Nat.Up.next h; loop f arr off _end (i+1) b else pure b -@[deprecated (since := "2024-08-19")] +@[deprecated "No deprecation message was provided." (since := "2024-08-19")] instance {m : Type u → Type v} : ForIn m ByteSlice UInt8 := ⟨fun ⟨arr, off, len⟩ b f ↦ forIn.loop f arr off (off + len) off b⟩ diff --git a/Mathlib/Deprecated/Combinator.lean b/Mathlib/Deprecated/Combinator.lean index 5c79bf0837a59..366fd5997f982 100644 --- a/Mathlib/Deprecated/Combinator.lean +++ b/Mathlib/Deprecated/Combinator.lean @@ -15,8 +15,11 @@ namespace Combinator universe u v w variable {α : Sort u} {β : Sort v} {γ : Sort w} -@[deprecated (since := "2024-07-27")] def I (a : α) := a -@[deprecated (since := "2024-07-27")] def K (a : α) (_b : β) := a -@[deprecated (since := "2024-07-27")] def S (x : α → β → γ) (y : α → β) (z : α) := x z (y z) +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] +def I (a : α) := a +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] +def K (a : α) (_b : β) := a +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] +def S (x : α → β → γ) (y : α → β) (z : α) := x z (y z) end Combinator diff --git a/Mathlib/Deprecated/Equiv.lean b/Mathlib/Deprecated/Equiv.lean index 74bc9b68ecdd3..b31fe578386cc 100644 --- a/Mathlib/Deprecated/Equiv.lean +++ b/Mathlib/Deprecated/Equiv.lean @@ -21,10 +21,10 @@ variable {α₁ β₁ : Type*} (e : α₁ ≃ β₁) (f : α₁ → α₁ → α set_option linter.deprecated false -@[deprecated (since := "2024-09-11")] +@[deprecated "No deprecation message was provided." (since := "2024-09-11")] instance [IsLeftCancel α₁ f] : IsLeftCancel β₁ (e.arrowCongr (e.arrowCongr e) f) := ⟨e.surjective.forall₃.2 fun x y z => by simpa using @IsLeftCancel.left_cancel _ f _ x y z⟩ -@[deprecated (since := "2024-09-11")] +@[deprecated "No deprecation message was provided." (since := "2024-09-11")] instance [IsRightCancel α₁ f] : IsRightCancel β₁ (e.arrowCongr (e.arrowCongr e) f) := ⟨e.surjective.forall₃.2 fun x y z => by simpa using @IsRightCancel.right_cancel _ f _ x y z⟩ diff --git a/Mathlib/Deprecated/LazyList.lean b/Mathlib/Deprecated/LazyList.lean index 7a9b820336505..5ba81d38999a8 100644 --- a/Mathlib/Deprecated/LazyList.lean +++ b/Mathlib/Deprecated/LazyList.lean @@ -24,7 +24,7 @@ namespace LazyList open Function /-- Isomorphism between strict and lazy lists. -/ -@[deprecated (since := "2024-07-22")] +@[deprecated "No deprecation message was provided." (since := "2024-07-22")] def listEquivLazyList (α : Type*) : List α ≃ LazyList α where toFun := LazyList.ofList invFun := LazyList.toList @@ -39,12 +39,12 @@ def listEquivLazyList (α : Type*) : List α ≃ LazyList α where · simp [toList, ofList] · simpa [ofList, toList] -@[deprecated (since := "2024-07-22")] +@[deprecated "No deprecation message was provided." (since := "2024-07-22")] instance : Traversable LazyList where map := @LazyList.traverse Id _ traverse := @LazyList.traverse -@[deprecated (since := "2024-07-22")] +@[deprecated "No deprecation message was provided." (since := "2024-07-22")] instance : LawfulTraversable LazyList := by apply Equiv.isLawfulTraversable' listEquivLazyList <;> intros <;> ext <;> rename_i f xs · induction xs using LazyList.rec with @@ -71,7 +71,7 @@ instance : LawfulTraversable LazyList := by Function.comp_def, Thunk.pure, ofList] | mk _ ih => apply ih -@[deprecated (since := "2024-07-22"), simp] +@[deprecated "No deprecation message was provided." (since := "2024-07-22"), simp] theorem bind_singleton {α} (x : LazyList α) : x.bind singleton = x := by induction x using LazyList.rec (motive_2 := fun xs => xs.get.bind singleton = xs.get) with | nil => simp [LazyList.bind] @@ -81,7 +81,7 @@ theorem bind_singleton {α} (x : LazyList α) : x.bind singleton = x := by simp [ih] | mk f ih => simp_all -@[deprecated (since := "2024-07-22")] +@[deprecated "No deprecation message was provided." (since := "2024-07-22")] instance : LawfulMonad LazyList := LawfulMonad.mk' (id_map := by intro α xs diff --git a/Mathlib/Deprecated/Logic.lean b/Mathlib/Deprecated/Logic.lean index f88642ec745a7..52d03620dc96a 100644 --- a/Mathlib/Deprecated/Logic.lean +++ b/Mathlib/Deprecated/Logic.lean @@ -37,43 +37,43 @@ local infix:65 (priority := high) " + " => g def Commutative := ∀ a b, a * b = b * a @[deprecated Std.Associative (since := "2024-09-13")] def Associative := ∀ a b c, (a * b) * c = a * (b * c) -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def LeftIdentity := ∀ a, one * a = a -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def RightIdentity := ∀ a, a * one = a -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def RightInverse := ∀ a, a * a⁻¹ = one -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def LeftCancelative := ∀ a b c, a * b = a * c → b = c -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def RightCancelative := ∀ a b c, a * b = c * b → a = c -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def LeftDistributive := ∀ a b c, a * (b + c) = a * b + a * c -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def RightDistributive := ∀ a b c, (a + b) * c = a * c + b * c end Binary @[deprecated (since := "2024-09-03")] alias not_of_eq_false := of_eq_false -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem cast_proof_irrel {β : Sort u} (h₁ h₂ : α = β) (a : α) : cast h₁ a = cast h₂ a := rfl @[deprecated (since := "2024-09-03")] alias eq_rec_heq := eqRec_heq @[deprecated (since := "2024-09-03")] alias heq_prop := proof_irrel_heq -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem heq_of_eq_rec_left {φ : α → Sort v} {a a' : α} {p₁ : φ a} {p₂ : φ a'} : (e : a = a') → (h₂ : Eq.rec (motive := fun a _ ↦ φ a) p₁ e = p₂) → HEq p₁ p₂ | rfl, rfl => HEq.rfl -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem heq_of_eq_rec_right {φ : α → Sort v} {a a' : α} {p₁ : φ a} {p₂ : φ a'} : (e : a' = a) → (h₂ : p₁ = Eq.rec (motive := fun a _ ↦ φ a) p₂ e) → HEq p₁ p₂ | rfl, rfl => HEq.rfl -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem of_heq_true {a : Prop} (h : HEq a True) : a := of_eq_true (eq_of_heq h) -@[deprecated (since := "2024-09-03")] +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] theorem eq_rec_compose {α β φ : Sort u} : ∀ (p₁ : β = φ) (p₂ : α = β) (a : α), (Eq.recOn p₁ (Eq.recOn p₂ a : β) : φ) = Eq.recOn (Eq.trans p₂ p₁) a @@ -114,20 +114,20 @@ theorem iff_self_iff (a : Prop) : (a ↔ a) ↔ True := iff_of_eq (iff_self _) /- decidable -/ -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem decide_True' (h : Decidable True) : decide True = true := by simp -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem decide_False' (h : Decidable False) : decide False = false := by simp namespace Decidable -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def recOn_true [h : Decidable p] {h₁ : p → Sort u} {h₂ : ¬p → Sort u} (h₃ : p) (h₄ : h₁ h₃) : Decidable.recOn h h₂ h₁ := cast (by match h with | .isTrue _ => rfl) h₄ -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def recOn_false [h : Decidable p] {h₁ : p → Sort u} {h₂ : ¬p → Sort u} (h₃ : ¬p) (h₄ : h₂ h₃) : Decidable.recOn h h₂ h₁ := cast (by match h with | .isFalse _ => rfl) h₄ @@ -145,25 +145,25 @@ end Decidable @[deprecated (since := "2024-09-03")] alias decidableTrue := instDecidableTrue @[deprecated (since := "2024-09-03")] alias decidableFalse := instDecidableFalse -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def IsDecEq {α : Sort u} (p : α → α → Bool) : Prop := ∀ ⦃x y : α⦄, p x y = true → x = y -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def IsDecRefl {α : Sort u} (p : α → α → Bool) : Prop := ∀ x, p x x = true -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def decidableEq_of_bool_pred {α : Sort u} {p : α → α → Bool} (h₁ : IsDecEq p) (h₂ : IsDecRefl p) : DecidableEq α | x, y => if hp : p x y = true then isTrue (h₁ hp) else isFalse (fun hxy : x = y ↦ absurd (h₂ y) (by rwa [hxy] at hp)) -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem decidableEq_inl_refl {α : Sort u} [h : DecidableEq α] (a : α) : h a a = isTrue (Eq.refl a) := match h a a with | isTrue _ => rfl -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem decidableEq_inr_neg {α : Sort u} [h : DecidableEq α] {a b : α} (n : a ≠ b) : h a b = isFalse n := match h a b with @@ -171,7 +171,7 @@ theorem decidableEq_inr_neg {α : Sort u} [h : DecidableEq α] {a b : α} /- subsingleton -/ -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem rec_subsingleton {p : Prop} [h : Decidable p] {h₁ : p → Sort u} {h₂ : ¬p → Sort u} [h₃ : ∀ h : p, Subsingleton (h₁ h)] [h₄ : ∀ h : ¬p, Subsingleton (h₂ h)] : Subsingleton (Decidable.recOn h h₂ h₁) := @@ -179,15 +179,15 @@ theorem rec_subsingleton {p : Prop} [h : Decidable p] {h₁ : p → Sort u} {h | isTrue h => h₃ h | isFalse h => h₄ h -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem imp_of_if_pos {c t e : Prop} [Decidable c] (h : ite c t e) (hc : c) : t := (if_pos hc ▸ h :) -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem imp_of_if_neg {c t e : Prop} [Decidable c] (h : ite c t e) (hnc : ¬c) : e := (if_neg hnc ▸ h :) -@[deprecated (since := "2024-09-11")] +@[deprecated "No deprecation message was provided." (since := "2024-09-11")] theorem dif_ctx_congr {α : Sort u} {b c : Prop} [dec_b : Decidable b] [dec_c : Decidable c] {x : b → α} {u : c → α} {y : ¬b → α} {v : ¬c → α} (h_c : b ↔ c) (h_t : ∀ h : c, x (Iff.mpr h_c h) = u h) @@ -199,7 +199,7 @@ theorem dif_ctx_congr {α : Sort u} {b c : Prop} [dec_b : Decidable b] [dec_c : | isFalse h₁, isTrue h₂ => absurd h₂ (Iff.mp (not_congr h_c) h₁) | isTrue h₁, isFalse h₂ => absurd h₁ (Iff.mpr (not_congr h_c) h₂) -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem if_ctx_congr_prop {b c x y u v : Prop} [dec_b : Decidable b] [dec_c : Decidable c] (h_c : b ↔ c) (h_t : c → (x ↔ u)) (h_e : ¬c → (y ↔ v)) : ite b x y ↔ ite c u v := match dec_b, dec_c with @@ -209,12 +209,12 @@ theorem if_ctx_congr_prop {b c x y u v : Prop} [dec_b : Decidable b] [dec_c : De | isTrue h₁, isFalse h₂ => absurd h₁ (Iff.mpr (not_congr h_c) h₂) -- @[congr] -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem if_congr_prop {b c x y u v : Prop} [Decidable b] [Decidable c] (h_c : b ↔ c) (h_t : x ↔ u) (h_e : y ↔ v) : ite b x y ↔ ite c u v := if_ctx_congr_prop h_c (fun _ ↦ h_t) (fun _ ↦ h_e) -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem if_ctx_simp_congr_prop {b c x y u v : Prop} [Decidable b] (h_c : b ↔ c) (h_t : c → (x ↔ u)) -- FIXME: after https://github.com/leanprover/lean4/issues/1867 is fixed, -- this should be changed back to: @@ -222,7 +222,7 @@ theorem if_ctx_simp_congr_prop {b c x y u v : Prop} [Decidable b] (h_c : b ↔ c (h_e : ¬c → (y ↔ v)) : ite b x y ↔ @ite _ c (decidable_of_decidable_of_iff h_c) u v := if_ctx_congr_prop (dec_c := decidable_of_decidable_of_iff h_c) h_c h_t h_e -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem if_simp_congr_prop {b c x y u v : Prop} [Decidable b] (h_c : b ↔ c) (h_t : x ↔ u) -- FIXME: after https://github.com/leanprover/lean4/issues/1867 is fixed, -- this should be changed back to: @@ -230,7 +230,7 @@ theorem if_simp_congr_prop {b c x y u v : Prop} [Decidable b] (h_c : b ↔ c) (h (h_e : y ↔ v) : ite b x y ↔ (@ite _ c (decidable_of_decidable_of_iff h_c) u v) := if_ctx_simp_congr_prop h_c (fun _ ↦ h_t) (fun _ ↦ h_e) -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem dif_ctx_simp_congr {α : Sort u} {b c : Prop} [Decidable b] {x : b → α} {u : c → α} {y : ¬b → α} {v : ¬c → α} (h_c : b ↔ c) (h_t : ∀ h : c, x (Iff.mpr h_c h) = u h) @@ -241,31 +241,31 @@ theorem dif_ctx_simp_congr {α : Sort u} {b c : Prop} [Decidable b] dite b x y = @dite _ c (decidable_of_decidable_of_iff h_c) u v := dif_ctx_congr (dec_c := decidable_of_decidable_of_iff h_c) h_c h_t h_e -@[deprecated (since := "2024-09-03")] +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] def AsTrue (c : Prop) [Decidable c] : Prop := if c then True else False -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def AsFalse (c : Prop) [Decidable c] : Prop := if c then False else True -@[deprecated (since := "2024-09-03")] +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] theorem AsTrue.get {c : Prop} [h₁ : Decidable c] (_ : AsTrue c) : c := match h₁ with | isTrue h_c => h_c /- Equalities for rewriting let-expressions -/ -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem let_value_eq {α : Sort u} {β : Sort v} {a₁ a₂ : α} (b : α → β) (h : a₁ = a₂) : (let x : α := a₁; b x) = (let x : α := a₂; b x) := congrArg b h -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem let_value_heq {α : Sort v} {β : α → Sort u} {a₁ a₂ : α} (b : ∀ x : α, β x) (h : a₁ = a₂) : HEq (let x : α := a₁; b x) (let x : α := a₂; b x) := by cases h; rfl -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem let_body_eq {α : Sort v} {β : α → Sort u} (a : α) {b₁ b₂ : ∀ x : α, β x} (h : ∀ x, b₁ x = b₂ x) : (let x : α := a; b₁ x) = (let x : α := a; b₂ x) := by exact h _ ▸ rfl -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem let_eq {α : Sort v} {β : Sort u} {a₁ a₂ : α} {b₁ b₂ : α → β} (h₁ : a₁ = a₂) (h₂ : ∀ x, b₁ x = b₂ x) : (let x : α := a₁; b₁ x) = (let x : α := a₂; b₂ x) := by simp [h₁, h₂] diff --git a/Mathlib/Deprecated/NatLemmas.lean b/Mathlib/Deprecated/NatLemmas.lean index 37fd557b3fec3..81b124b4476cd 100644 --- a/Mathlib/Deprecated/NatLemmas.lean +++ b/Mathlib/Deprecated/NatLemmas.lean @@ -28,7 +28,7 @@ namespace Nat /-! successor and predecessor -/ -@[deprecated (since := "2024-08-23")] +@[deprecated "No deprecation message was provided." (since := "2024-08-23")] def discriminate {B : Sort u} {n : ℕ} (H1 : n = 0 → B) (H2 : ∀ m, n = succ m → B) : B := by induction n with | zero => exact H1 rfl @@ -36,7 +36,7 @@ def discriminate {B : Sort u} {n : ℕ} (H1 : n = 0 → B) (H2 : ∀ m, n = succ -- Unused in Mathlib; -- if downstream projects find this essential please copy it or remove the deprecation. -@[deprecated (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] theorem one_eq_succ_zero : 1 = succ 0 := rfl @@ -44,7 +44,7 @@ theorem one_eq_succ_zero : 1 = succ 0 := -- Unused in Mathlib; -- if downstream projects find this essential please copy it or remove the deprecation. -@[deprecated (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] def subInduction {P : ℕ → ℕ → Sort u} (H1 : ∀ m, P 0 m) (H2 : ∀ n, P (succ n) 0) (H3 : ∀ n m, P n m → P (succ n) (succ m)) : ∀ n m : ℕ, P n m | 0, _m => H1 _ @@ -55,7 +55,7 @@ def subInduction {P : ℕ → ℕ → Sort u} (H1 : ∀ m, P 0 m) (H2 : ∀ n, P -- Unused in Mathlib; -- if downstream projects find this essential please copy it or remove the deprecation. -@[deprecated (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] theorem cond_decide_mod_two (x : ℕ) [d : Decidable (x % 2 = 1)] : cond (@decide (x % 2 = 1) d) 1 0 = x % 2 := by simp only [cond_eq_if, decide_eq_true_eq] diff --git a/Mathlib/Deprecated/RelClasses.lean b/Mathlib/Deprecated/RelClasses.lean index 357b10283c2d5..8d5f647030d5d 100644 --- a/Mathlib/Deprecated/RelClasses.lean +++ b/Mathlib/Deprecated/RelClasses.lean @@ -27,12 +27,14 @@ variable {α : Type u} open Function -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem IsTotalPreorder.swap (r) [IsTotalPreorder α r] : IsTotalPreorder α (swap r) := { @IsPreorder.swap α r _, @IsTotal.swap α r _ with } -@[deprecated (since := "2024-08-22")] instance [LinearOrder α] : IsTotalPreorder α (· ≤ ·) where -@[deprecated (since := "2024-08-22")] instance [LinearOrder α] : IsTotalPreorder α (· ≥ ·) where +@[deprecated "No deprecation message was provided." (since := "2024-08-22")] +instance [LinearOrder α] : IsTotalPreorder α (· ≤ ·) where +@[deprecated "No deprecation message was provided." (since := "2024-08-22")] +instance [LinearOrder α] : IsTotalPreorder α (· ≥ ·) where -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] instance [LinearOrder α] : IsIncompTrans α (· < ·) := by infer_instance diff --git a/Mathlib/GroupTheory/Coset/Basic.lean b/Mathlib/GroupTheory/Coset/Basic.lean index a93c6f92a2109..4f8c575b67fcf 100644 --- a/Mathlib/GroupTheory/Coset/Basic.lean +++ b/Mathlib/GroupTheory/Coset/Basic.lean @@ -318,9 +318,15 @@ lemma orbit_eq_out_smul (x : α ⧸ s) : MulAction.orbitRel.Quotient.orbit x = x induction x using QuotientGroup.induction_on simp only [orbit_mk_eq_smul, ← eq_class_eq_leftCoset, Quotient.out_eq'] -@[to_additive (attr := deprecated (since := "2024-10-19"))] +@[to_additive] alias orbit_eq_out'_smul := orbit_eq_out_smul +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated orbit_eq_out_smul (since := "2024-10-19")] orbit_eq_out'_smul +attribute [deprecated QuotientAddGroup.orbit_eq_out_vadd (since := "2024-10-19")] +QuotientAddGroup.orbit_eq_out'_vadd + end QuotientGroup namespace Subgroup diff --git a/Mathlib/GroupTheory/Coset/Defs.lean b/Mathlib/GroupTheory/Coset/Defs.lean index bd49ccf9a7568..c57164adfbaee 100644 --- a/Mathlib/GroupTheory/Coset/Defs.lean +++ b/Mathlib/GroupTheory/Coset/Defs.lean @@ -174,7 +174,13 @@ theorem induction_on {C : α ⧸ s → Prop} (x : α ⧸ s) (H : ∀ z, C (Quoti instance : Coe α (α ⧸ s) := ⟨mk⟩ -@[to_additive (attr := deprecated (since := "2024-08-04"))] alias induction_on' := induction_on +@[to_additive] alias induction_on' := induction_on + +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated induction_on (since := "2024-08-04")] induction_on' +attribute [deprecated QuotientAddGroup.induction_on (since := "2024-08-04")] +QuotientAddGroup.induction_on' @[to_additive (attr := simp)] theorem quotient_liftOn_mk {β} (f : α → β) (h) (x : α) : Quotient.liftOn' (x : α ⧸ s) f h = f x := @@ -198,7 +204,8 @@ protected theorem eq {a b : α} : (a : α ⧸ s) = b ↔ a⁻¹ * b ∈ s := _ ↔ leftRel s a b := Quotient.eq'' _ ↔ _ := by rw [leftRel_apply] -@[to_additive (attr := deprecated (since := "2024-08-04"))] alias eq' := QuotientGroup.eq +@[to_additive (attr := deprecated "No deprecation message was provided." (since := "2024-08-04"))] +alias eq' := QuotientGroup.eq @[to_additive] theorem out_eq' (a : α ⧸ s) : mk a.out = a := @@ -213,9 +220,16 @@ variable (s) theorem mk_out_eq_mul (g : α) : ∃ h : s, (mk g : α ⧸ s).out = g * h := ⟨⟨g⁻¹ * (mk g).out, QuotientGroup.eq.mp (mk g).out_eq'.symm⟩, by rw [mul_inv_cancel_left]⟩ -@[to_additive (attr := deprecated (since := "2024-10-19")) QuotientAddGroup.mk_out'_eq_mul] +@[to_additive QuotientAddGroup.mk_out'_eq_mul] alias mk_out'_eq_mul := mk_out_eq_mul +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated mk_out_eq_mul (since := "2024-10-19")] mk_out'_eq_mul +attribute [deprecated QuotientAddGroup.mk_out_eq_mul (since := "2024-10-19")] +QuotientAddGroup.mk_out'_eq_mul + + variable {s} {a b : α} @[to_additive (attr := simp)] diff --git a/Mathlib/GroupTheory/Exponent.lean b/Mathlib/GroupTheory/Exponent.lean index 6ee49e4b8a6ac..ea64a5ce67e5e 100644 --- a/Mathlib/GroupTheory/Exponent.lean +++ b/Mathlib/GroupTheory/Exponent.lean @@ -644,7 +644,8 @@ lemma inv_eq_self_of_orderOf_eq_two {x : G} (hx : orderOf x = 2) : -- TODO: delete /-- Any group of exponent two is abelian. -/ -@[to_additive (attr := reducible, deprecated (since := "2024-02-17")) +@[to_additive (attr := reducible, + deprecated "No deprecation message was provided." (since := "2024-02-17")) "Any additive group of exponent two is abelian."] def instCommGroupOfExponentTwo (hG : Monoid.exponent G = 2) : CommGroup G where mul_comm := mul_comm_of_exponent_two hG diff --git a/Mathlib/GroupTheory/GroupAction/Defs.lean b/Mathlib/GroupTheory/GroupAction/Defs.lean index 3848050964715..fa15f5e7d6253 100644 --- a/Mathlib/GroupTheory/GroupAction/Defs.lean +++ b/Mathlib/GroupTheory/GroupAction/Defs.lean @@ -312,9 +312,15 @@ variable {G α} theorem orbitRel_apply {a b : α} : orbitRel G α a b ↔ a ∈ orbit G b := Iff.rfl -@[to_additive (attr := deprecated (since := "2024-10-18"))] +@[to_additive] alias orbitRel_r_apply := orbitRel_apply +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated orbitRel_apply (since := "2024-10-18")] orbitRel_r_apply +attribute [deprecated AddAction.orbitRel_apply (since := "2024-10-18")] AddAction.orbitRel_r_apply + + /-- When you take a set `U` in `α`, push it down to the quotient, and pull back, you get the union of the orbit of `U` under `G`. -/ @[to_additive diff --git a/Mathlib/GroupTheory/GroupAction/Quotient.lean b/Mathlib/GroupTheory/GroupAction/Quotient.lean index 4c15d5da76105..778fb1159400a 100644 --- a/Mathlib/GroupTheory/GroupAction/Quotient.lean +++ b/Mathlib/GroupTheory/GroupAction/Quotient.lean @@ -103,13 +103,26 @@ theorem _root_.QuotientGroup.out_conj_pow_minimalPeriod_mem (a : α) (q : α ⧸ rw [mul_assoc, ← QuotientGroup.eq, QuotientGroup.out_eq', ← smul_eq_mul, Quotient.mk_smul_out, eq_comm, pow_smul_eq_iff_minimalPeriod_dvd] -@[to_additive (attr := deprecated (since := "2024-10-19"))] +@[to_additive] alias Quotient.mk_smul_out' := Quotient.mk_smul_out +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated Quotient.mk_smul_out (since := "2024-10-19")] Quotient.mk_smul_out' +attribute [deprecated AddAction.Quotient.mk_vadd_out (since := "2024-10-19")] +AddAction.Quotient.mk_vadd_out' + -- Porting note: removed simp attribute, simp can prove this -@[to_additive (attr := deprecated (since := "2024-10-19"))] +@[to_additive] alias Quotient.coe_smul_out' := Quotient.coe_smul_out +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated Quotient.coe_smul_out (since := "2024-10-19")] Quotient.coe_smul_out' +attribute [deprecated AddAction.Quotient.coe_vadd_out (since := "2024-10-19")] +AddAction.Quotient.coe_vadd_out' + + @[deprecated (since := "2024-10-19")] alias _root_.QuotientGroup.out'_conj_pow_minimalPeriod_mem := QuotientGroup.out_conj_pow_minimalPeriod_mem diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index b380c1f71c410..b18e30f62e781 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -88,9 +88,14 @@ theorem not_isOfFinOrder_of_injective_pow {x : G} (h : Injective fun n : ℕ => theorem IsOfFinOrder.one : IsOfFinOrder (1 : G) := isOfFinOrder_iff_pow_eq_one.mpr ⟨1, Nat.one_pos, one_pow 1⟩ -@[to_additive (attr := deprecated (since := "2024-10-11"))] +@[to_additive] alias isOfFinOrder_one := IsOfFinOrder.one +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated IsOfFinOrder.one (since := "2024-10-11")] isOfFinOrder_one +attribute [deprecated IsOfFinAddOrder.zero (since := "2024-10-11")] isOfFinAddOrder_zero + @[to_additive] lemma IsOfFinOrder.pow {n : ℕ} : IsOfFinOrder a → IsOfFinOrder (a ^ n) := by simp_rw [isOfFinOrder_iff_pow_eq_one] diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index 11e9dc5991b44..d13254a943cc7 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -327,10 +327,17 @@ lemma isCyclic_iff_exists_ofOrder_eq_natCard [Finite α] : refine isCyclic_of_orderOf_eq_card g ?_ simp [hg] -@[to_additive (attr := deprecated (since := "2024-04-20"))] +@[to_additive] protected alias IsCyclic.iff_exists_ofOrder_eq_natCard_of_Fintype := isCyclic_iff_exists_ofOrder_eq_natCard +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated isCyclic_iff_exists_ofOrder_eq_natCard (since := "2024-04-20")] +IsCyclic.iff_exists_ofOrder_eq_natCard_of_Fintype +attribute [deprecated isAddCyclic_iff_exists_ofOrder_eq_natCard (since := "2024-04-20")] +IsAddCyclic.iff_exists_ofOrder_eq_natCard_of_Fintype + section variable [Fintype α] diff --git a/Mathlib/LinearAlgebra/BilinearForm/Basic.lean b/Mathlib/LinearAlgebra/BilinearForm/Basic.lean index 0768933dedc46..da171b8d6bfaa 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Basic.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Basic.lean @@ -58,7 +58,7 @@ namespace LinearMap namespace BilinForm -@[deprecated (since := "2024-04-14")] +@[deprecated "No deprecation message was provided." (since := "2024-04-14")] theorem coeFn_congr : ∀ {x x' y y' : M}, x = x' → y = y' → B x y = B x' y' | _, _, _, _, rfl, rfl => rfl @@ -101,7 +101,7 @@ theorem ext (H : ∀ x y : M, B x y = D x y) : B = D := ext₂ H theorem congr_fun (h : B = D) (x y : M) : B x y = D x y := congr_fun₂ h _ _ -@[deprecated (since := "2024-04-14")] +@[deprecated "No deprecation message was provided." (since := "2024-04-14")] theorem coe_zero : ⇑(0 : BilinForm R M) = 0 := rfl @@ -111,7 +111,7 @@ theorem zero_apply (x y : M) : (0 : BilinForm R M) x y = 0 := variable (B D B₁ D₁) -@[deprecated (since := "2024-04-14")] +@[deprecated "No deprecation message was provided." (since := "2024-04-14")] theorem coe_add : ⇑(B + D) = B + D := rfl @@ -119,7 +119,7 @@ theorem coe_add : ⇑(B + D) = B + D := theorem add_apply (x y : M) : (B + D) x y = B x y + D x y := rfl -@[deprecated (since := "2024-04-14")] +@[deprecated "No deprecation message was provided." (since := "2024-04-14")] theorem coe_neg : ⇑(-B₁) = -B₁ := rfl @@ -127,7 +127,7 @@ theorem coe_neg : ⇑(-B₁) = -B₁ := theorem neg_apply (x y : M₁) : (-B₁) x y = -B₁ x y := rfl -@[deprecated (since := "2024-04-14")] +@[deprecated "No deprecation message was provided." (since := "2024-04-14")] theorem coe_sub : ⇑(B₁ - D₁) = B₁ - D₁ := rfl diff --git a/Mathlib/LinearAlgebra/BilinearForm/Hom.lean b/Mathlib/LinearAlgebra/BilinearForm/Hom.lean index 9c011c36ed52b..c1b171f04d13e 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Hom.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Hom.lean @@ -58,16 +58,16 @@ section ToLin' def toLinHomAux₁ (A : BilinForm R M) (x : M) : M →ₗ[R] R := A x /-- Auxiliary definition to define `toLinHom`; see below. -/ -@[deprecated (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] def toLinHomAux₂ (A : BilinForm R M) : M →ₗ[R] M →ₗ[R] R := A /-- The linear map obtained from a `BilinForm` by fixing the left co-ordinate and evaluating in the right. -/ -@[deprecated (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] def toLinHom : BilinForm R M →ₗ[R] M →ₗ[R] M →ₗ[R] R := LinearMap.id set_option linter.deprecated false in -@[deprecated (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] theorem toLin'_apply (A : BilinForm R M) (x : M) : toLinHom (M := M) A x = A x := rfl @@ -112,7 +112,7 @@ def LinearMap.toBilinAux (f : M →ₗ[R] M →ₗ[R] R) : BilinForm R M := f set_option linter.deprecated false in /-- Bilinear forms are linearly equivalent to maps with two arguments that are linear in both. -/ -@[deprecated (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] def LinearMap.BilinForm.toLin : BilinForm R M ≃ₗ[R] M →ₗ[R] M →ₗ[R] R := { BilinForm.toLinHom with invFun := LinearMap.toBilinAux @@ -121,35 +121,35 @@ def LinearMap.BilinForm.toLin : BilinForm R M ≃ₗ[R] M →ₗ[R] M →ₗ[R] set_option linter.deprecated false in /-- A map with two arguments that is linear in both is linearly equivalent to bilinear form. -/ -@[deprecated (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] def LinearMap.toBilin : (M →ₗ[R] M →ₗ[R] R) ≃ₗ[R] BilinForm R M := BilinForm.toLin.symm -@[deprecated (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] theorem LinearMap.toBilinAux_eq (f : M →ₗ[R] M →ₗ[R] R) : LinearMap.toBilinAux f = f := rfl set_option linter.deprecated false in -@[deprecated (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] theorem LinearMap.toBilin_symm : (LinearMap.toBilin.symm : BilinForm R M ≃ₗ[R] _) = BilinForm.toLin := rfl set_option linter.deprecated false in -@[deprecated (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] theorem BilinForm.toLin_symm : (BilinForm.toLin.symm : _ ≃ₗ[R] BilinForm R M) = LinearMap.toBilin := LinearMap.toBilin.symm_symm set_option linter.deprecated false in -@[deprecated (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] theorem LinearMap.toBilin_apply (f : M →ₗ[R] M →ₗ[R] R) (x y : M) : toBilin f x y = f x y := rfl set_option linter.deprecated false in -@[deprecated (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] theorem BilinForm.toLin_apply (x : M) : BilinForm.toLin B x = B x := rfl diff --git a/Mathlib/LinearAlgebra/Prod.lean b/Mathlib/LinearAlgebra/Prod.lean index 4a19736a68ba4..7bbb4ac2f4b65 100644 --- a/Mathlib/LinearAlgebra/Prod.lean +++ b/Mathlib/LinearAlgebra/Prod.lean @@ -834,17 +834,17 @@ set_option linter.deprecated false /-- An auxiliary construction for `tunnel`. The composition of `f`, followed by the isomorphism back to `K`, followed by the inclusion of this submodule back into `M`. -/ -@[deprecated (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] def tunnelAux (f : M × N →ₗ[R] M) (Kφ : ΣK : Submodule R M, K ≃ₗ[R] M) : M × N →ₗ[R] M := (Kφ.1.subtype.comp Kφ.2.symm.toLinearMap).comp f -@[deprecated (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tunnelAux_injective (f : M × N →ₗ[R] M) (i : Injective f) (Kφ : ΣK : Submodule R M, K ≃ₗ[R] M) : Injective (tunnelAux f Kφ) := (Subtype.val_injective.comp Kφ.2.symm.injective).comp i /-- Auxiliary definition for `tunnel`. -/ -@[deprecated (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] def tunnel' (f : M × N →ₗ[R] M) (i : Injective f) : ℕ → ΣK : Submodule R M, K ≃ₗ[R] M | 0 => ⟨⊤, LinearEquiv.ofTop ⊤ rfl⟩ | n + 1 => @@ -855,7 +855,7 @@ def tunnel' (f : M × N →ₗ[R] M) (i : Injective f) : ℕ → ΣK : Submodule /-- Give an injective map `f : M × N →ₗ[R] M` we can find a nested sequence of submodules all isomorphic to `M`. -/ -@[deprecated (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] def tunnel (f : M × N →ₗ[R] M) (i : Injective f) : ℕ →o (Submodule R M)ᵒᵈ := -- Note: the hint `(α := _)` had to be added in https://github.com/leanprover-community/mathlib4/pull/8386 ⟨fun n => OrderDual.toDual (α := Submodule R M) (tunnel' f i n).1, @@ -867,24 +867,24 @@ def tunnel (f : M × N →ₗ[R] M) (i : Injective f) : ℕ →o (Submodule R M) /-- Give an injective map `f : M × N →ₗ[R] M` we can find a sequence of submodules all isomorphic to `N`. -/ -@[deprecated (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] def tailing (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Submodule R M := (Submodule.snd R M N).map (tunnelAux f (tunnel' f i n)) /-- Each `tailing f i n` is a copy of `N`. -/ -@[deprecated (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] def tailingLinearEquiv (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailing f i n ≃ₗ[R] N := ((Submodule.snd R M N).equivMapOfInjective _ (tunnelAux_injective f i (tunnel' f i n))).symm.trans (Submodule.sndEquiv R M N) -@[deprecated (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailing_le_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailing f i n ≤ OrderDual.ofDual (α := Submodule R M) (tunnel f i n) := by dsimp [tailing, tunnelAux] rw [Submodule.map_comp, Submodule.map_comp] apply Submodule.map_subtype_le -@[deprecated (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailing_disjoint_tunnel_succ (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Disjoint (tailing f i n) (OrderDual.ofDual (α := Submodule R M) <| tunnel f i (n + 1)) := by rw [disjoint_iff] @@ -893,7 +893,7 @@ theorem tailing_disjoint_tunnel_succ (f : M × N →ₗ[R] M) (i : Injective f) Submodule.comap_map_eq_of_injective (tunnelAux_injective _ i _), inf_comm, Submodule.fst_inf_snd, Submodule.map_bot] -@[deprecated (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailing_sup_tunnel_succ_le_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailing f i n ⊔ (OrderDual.ofDual (α := Submodule R M) <| tunnel f i (n + 1)) ≤ (OrderDual.ofDual (α := Submodule R M) <| tunnel f i n) := by @@ -902,19 +902,19 @@ theorem tailing_sup_tunnel_succ_le_tunnel (f : M × N →ₗ[R] M) (i : Injectiv apply Submodule.map_subtype_le /-- The supremum of all the copies of `N` found inside the tunnel. -/ -@[deprecated (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] def tailings (f : M × N →ₗ[R] M) (i : Injective f) : ℕ → Submodule R M := partialSups (tailing f i) -@[simp, deprecated (since := "2024-06-05")] +@[simp, deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailings_zero (f : M × N →ₗ[R] M) (i : Injective f) : tailings f i 0 = tailing f i 0 := by simp [tailings] -@[simp, deprecated (since := "2024-06-05")] +@[simp, deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailings_succ (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailings f i (n + 1) = tailings f i n ⊔ tailing f i (n + 1) := by simp [tailings] -@[deprecated (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailings_disjoint_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Disjoint (tailings f i n) (OrderDual.ofDual (α := Submodule R M) <| tunnel f i (n + 1)) := by induction' n with n ih @@ -926,7 +926,7 @@ theorem tailings_disjoint_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : · apply Disjoint.mono_right _ ih apply tailing_sup_tunnel_succ_le_tunnel -@[deprecated (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailings_disjoint_tailing (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Disjoint (tailings f i n) (tailing f i (n + 1)) := Disjoint.mono_right (tailing_le_tunnel f i _) (tailings_disjoint_tunnel f i _) diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index 7a0eb1ff4d84f..89c3a8fe8a2f1 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -725,20 +725,21 @@ alias by_contradiction := byContradiction -- TODO: remove? rename in core? alias prop_complete := propComplete -- TODO: remove? rename in core? -@[elab_as_elim, deprecated (since := "2024-07-27")] theorem cases_true_false (p : Prop → Prop) +@[elab_as_elim, deprecated "No deprecation message was provided." (since := "2024-07-27")] +theorem cases_true_false (p : Prop → Prop) (h1 : p True) (h2 : p False) (a : Prop) : p a := Or.elim (prop_complete a) (fun ht : a = True ↦ ht.symm ▸ h1) fun hf : a = False ↦ hf.symm ▸ h2 -@[deprecated (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] theorem eq_false_or_eq_true (a : Prop) : a = False ∨ a = True := (prop_complete a).symm set_option linter.deprecated false in -@[deprecated (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] theorem cases_on (a : Prop) {p : Prop → Prop} (h1 : p True) (h2 : p False) : p a := @cases_true_false p h1 h2 a set_option linter.deprecated false in -@[deprecated (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] theorem cases {p : Prop → Prop} (h1 : p True) (h2 : p False) (a) : p a := cases_on a h1 h2 end Classical diff --git a/Mathlib/MeasureTheory/Function/L1Space.lean b/Mathlib/MeasureTheory/Function/L1Space.lean index 25e17db9f409b..b003cc9a9db03 100644 --- a/Mathlib/MeasureTheory/Function/L1Space.lean +++ b/Mathlib/MeasureTheory/Function/L1Space.lean @@ -443,7 +443,7 @@ lemma Integrable.of_finite [Finite α] [MeasurableSingletonClass α] [IsFiniteMe /-- This lemma is a special case of `Integrable.of_finite`. -/ -- Eternal deprecation for discoverability, don't remove -@[deprecated Integrable.of_finite, nolint deprecatedNoSince] +@[deprecated Integrable.of_finite (since := "2024-10-05"), nolint deprecatedNoSince] lemma Integrable.of_isEmpty [IsEmpty α] {f : α → β} : Integrable f μ := .of_finite @[deprecated (since := "2024-02-05")] alias integrable_of_fintype := Integrable.of_finite diff --git a/Mathlib/MeasureTheory/Measure/Typeclasses.lean b/Mathlib/MeasureTheory/Measure/Typeclasses.lean index c9a41869344f8..f8d833559d3e9 100644 --- a/Mathlib/MeasureTheory/Measure/Typeclasses.lean +++ b/Mathlib/MeasureTheory/Measure/Typeclasses.lean @@ -567,7 +567,7 @@ instance isFiniteMeasure_sfiniteSeq [h : SFinite μ] (n : ℕ) : IsFiniteMeasure h.1.choose_spec.1 n set_option linter.deprecated false in -@[deprecated (since := "2024-10-11")] +@[deprecated "No deprecation message was provided." (since := "2024-10-11")] instance isFiniteMeasure_sFiniteSeq [SFinite μ] (n : ℕ) : IsFiniteMeasure (sFiniteSeq μ n) := isFiniteMeasure_sfiniteSeq n diff --git a/Mathlib/MeasureTheory/Measure/WithDensityFinite.lean b/Mathlib/MeasureTheory/Measure/WithDensityFinite.lean index 21abd06cd6cee..40b56846328c7 100644 --- a/Mathlib/MeasureTheory/Measure/WithDensityFinite.lean +++ b/Mathlib/MeasureTheory/Measure/WithDensityFinite.lean @@ -117,7 +117,7 @@ noncomputable def Measure.densityToFinite (μ : Measure α) [SFinite μ] (a : α μ.rnDeriv μ.toFinite a set_option linter.deprecated false in -@[deprecated (since := "2024-10-04")] +@[deprecated "No deprecation message was provided." (since := "2024-10-04")] lemma densityToFinite_def (μ : Measure α) [SFinite μ] : μ.densityToFinite = μ.rnDeriv μ.toFinite := rfl diff --git a/Mathlib/NumberTheory/ArithmeticFunction.lean b/Mathlib/NumberTheory/ArithmeticFunction.lean index 5c13f9d53ebe2..8b8038b188ab0 100644 --- a/Mathlib/NumberTheory/ArithmeticFunction.lean +++ b/Mathlib/NumberTheory/ArithmeticFunction.lean @@ -1291,7 +1291,8 @@ theorem _root_.Nat.card_divisors {n : ℕ} (hn : n ≠ 0) : exact Finset.prod_congr n.support_factorization fun _ h => sigma_zero_apply_prime_pow <| Nat.prime_of_mem_primeFactors h -@[deprecated (since := "2024-06-09")] theorem card_divisors (n : ℕ) (hn : n ≠ 0) : +@[deprecated "No deprecation message was provided." (since := "2024-06-09")] +theorem card_divisors (n : ℕ) (hn : n ≠ 0) : #n.divisors = n.primeFactors.prod (n.factorization · + 1) := Nat.card_divisors hn theorem _root_.Nat.sum_divisors {n : ℕ} (hn : n ≠ 0) : diff --git a/Mathlib/NumberTheory/MulChar/Basic.lean b/Mathlib/NumberTheory/MulChar/Basic.lean index 2df970973c59d..7ddd95b920b4e 100644 --- a/Mathlib/NumberTheory/MulChar/Basic.lean +++ b/Mathlib/NumberTheory/MulChar/Basic.lean @@ -391,13 +391,13 @@ lemma ne_one_iff {χ : MulChar R R'} : χ ≠ 1 ↔ ∃ a : Rˣ, χ a ≠ 1 := b simp only [Ne, eq_one_iff, not_forall] /-- A multiplicative character is *nontrivial* if it takes a value `≠ 1` on a unit. -/ -@[deprecated (since := "2024-06-16")] +@[deprecated "No deprecation message was provided." (since := "2024-06-16")] def IsNontrivial (χ : MulChar R R') : Prop := ∃ a : Rˣ, χ a ≠ 1 set_option linter.deprecated false in /-- A multiplicative character is nontrivial iff it is not the trivial character. -/ -@[deprecated (since := "2024-06-16")] +@[deprecated "No deprecation message was provided." (since := "2024-06-16")] theorem isNontrivial_iff (χ : MulChar R R') : χ.IsNontrivial ↔ χ ≠ 1 := by simp only [IsNontrivial, Ne, MulChar.ext_iff, not_forall, one_apply_coe] @@ -573,7 +573,7 @@ theorem sum_eq_zero_of_ne_one [IsDomain R'] {χ : MulChar R R'} (hχ : χ ≠ 1) simpa only [Finset.mul_sum, ← map_mul] using b.mulLeft_bijective.sum_comp _ set_option linter.deprecated false in -@[deprecated (since := "2024-06-16")] +@[deprecated "No deprecation message was provided." (since := "2024-06-16")] lemma IsNontrivial.sum_eq_zero [IsDomain R'] {χ : MulChar R R'} (hχ : χ.IsNontrivial) : ∑ a, χ a = 0 := sum_eq_zero_of_ne_one ((isNontrivial_iff _).mp hχ) diff --git a/Mathlib/Order/Defs.lean b/Mathlib/Order/Defs.lean index dc90cc7b0adbe..b97cdc2c979d6 100644 --- a/Mathlib/Order/Defs.lean +++ b/Mathlib/Order/Defs.lean @@ -206,11 +206,11 @@ theorem Equivalence.transitive (h : Equivalence r) : Transitive r := variable {β : Sort*} (r : β → β → Prop) (f : α → β) -@[deprecated (since := "2024-09-13")] +@[deprecated "No deprecation message was provided." (since := "2024-09-13")] theorem InvImage.trans (h : Transitive r) : Transitive (InvImage r f) := fun (a₁ a₂ a₃ : α) (h₁ : InvImage r f a₁ a₂) (h₂ : InvImage r f a₂ a₃) ↦ h h₁ h₂ -@[deprecated (since := "2024-09-13")] +@[deprecated "No deprecation message was provided." (since := "2024-09-13")] theorem InvImage.irreflexive (h : Irreflexive r) : Irreflexive (InvImage r f) := fun (a : α) (h₁ : InvImage r f a a) ↦ h (f a) h₁ @@ -276,7 +276,7 @@ lemma lt_iff_le_not_le : a < b ↔ a ≤ b ∧ ¬b ≤ a := Preorder.lt_iff_le_n lemma lt_of_le_not_le (hab : a ≤ b) (hba : ¬ b ≤ a) : a < b := lt_iff_le_not_le.2 ⟨hab, hba⟩ -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem le_not_le_of_lt : ∀ {a b : α}, a < b → a ≤ b ∧ ¬b ≤ a | _a, _b, hab => lt_iff_le_not_le.mp hab @@ -502,7 +502,7 @@ namespace Nat /-! Deprecated properties of inequality on `Nat` -/ -@[deprecated (since := "2024-08-23")] +@[deprecated "No deprecation message was provided." (since := "2024-08-23")] protected def ltGeByCases {a b : Nat} {C : Sort*} (h₁ : a < b → C) (h₂ : b ≤ a → C) : C := Decidable.byCases h₁ fun h => h₂ (Or.elim (Nat.lt_or_ge a b) (fun a => absurd a h) fun a => a) diff --git a/Mathlib/Order/OmegaCompletePartialOrder.lean b/Mathlib/Order/OmegaCompletePartialOrder.lean index 4c97970762171..c34eaa969055e 100644 --- a/Mathlib/Order/OmegaCompletePartialOrder.lean +++ b/Mathlib/Order/OmegaCompletePartialOrder.lean @@ -731,7 +731,7 @@ theorem apply_mono {f g : α →𝒄 β} {x y : α} (h₁ : f ≤ g) (h₂ : x OrderHom.apply_mono (show (f : α →o β) ≤ g from h₁) h₂ set_option linter.deprecated false in -@[deprecated (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] theorem ite_continuous' {p : Prop} [hp : Decidable p] (f g : α → β) (hf : Continuous' f) (hg : Continuous' g) : Continuous' fun x => if p then f x else g x := by split_ifs <;> simp [*] diff --git a/Mathlib/Order/RelClasses.lean b/Mathlib/Order/RelClasses.lean index c99abc227fd09..0acae0e49466b 100644 --- a/Mathlib/Order/RelClasses.lean +++ b/Mathlib/Order/RelClasses.lean @@ -86,7 +86,7 @@ theorem IsStrictOrder.swap (r) [IsStrictOrder α r] : IsStrictOrder α (swap r) theorem IsPartialOrder.swap (r) [IsPartialOrder α r] : IsPartialOrder α (swap r) := { @IsPreorder.swap α r _, @IsAntisymm.swap α r _ with } -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem IsLinearOrder.swap (r) [IsLinearOrder α r] : IsLinearOrder α (swap r) := { @IsPartialOrder.swap α r _, @IsTotal.swap α r _ with } @@ -212,7 +212,7 @@ instance (priority := 100) isStrictOrderConnected_of_isStrictTotalOrder [IsStric fun o ↦ o.elim (fun e ↦ e ▸ h) fun h' ↦ _root_.trans h' h⟩ -- see Note [lower instance priority] -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] instance (priority := 100) isStrictTotalOrder_of_isStrictTotalOrder [IsStrictTotalOrder α r] : IsStrictWeakOrder α r := { isStrictWeakOrder_of_isOrderConnected with } @@ -780,7 +780,7 @@ instance [LinearOrder α] : IsStrictTotalOrder α (· < ·) where instance [LinearOrder α] : IsOrderConnected α (· < ·) := by infer_instance -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] instance [LinearOrder α] : IsStrictWeakOrder α (· < ·) := by infer_instance theorem transitive_le [Preorder α] : Transitive (@LE.le α _) := diff --git a/Mathlib/SetTheory/Cardinal/Aleph.lean b/Mathlib/SetTheory/Cardinal/Aleph.lean index dafc9ee33dced..ca79284dac29a 100644 --- a/Mathlib/SetTheory/Cardinal/Aleph.lean +++ b/Mathlib/SetTheory/Cardinal/Aleph.lean @@ -541,7 +541,7 @@ def alephIdx.relIso : @RelIso Cardinal.{u} Ordinal.{u} (· < ·) (· < ·) := def alephIdx : Cardinal → Ordinal := aleph'.symm -@[deprecated (since := "2024-08-28")] +@[deprecated "No deprecation message was provided." (since := "2024-08-28")] theorem alephIdx.relIso_coe : (alephIdx.relIso : Cardinal → Ordinal) = alephIdx := rfl @@ -555,7 +555,7 @@ theorem alephIdx.relIso_coe : (alephIdx.relIso : Cardinal → Ordinal) = alephId def Aleph'.relIso := aleph' -@[deprecated (since := "2024-08-28")] +@[deprecated "No deprecation message was provided." (since := "2024-08-28")] theorem aleph'.relIso_coe : (Aleph'.relIso : Ordinal → Cardinal) = aleph' := rfl @@ -571,11 +571,11 @@ theorem aleph'_le {o₁ o₂ : Ordinal} : aleph' o₁ ≤ aleph' o₂ ↔ o₁ theorem aleph'_max (o₁ o₂ : Ordinal) : aleph' (max o₁ o₂) = max (aleph' o₁) (aleph' o₂) := aleph'.monotone.map_max -@[deprecated (since := "2024-08-28")] +@[deprecated "No deprecation message was provided." (since := "2024-08-28")] theorem aleph'_alephIdx (c : Cardinal) : aleph' c.alephIdx = c := Cardinal.alephIdx.relIso.toEquiv.symm_apply_apply c -@[deprecated (since := "2024-08-28")] +@[deprecated "No deprecation message was provided." (since := "2024-08-28")] theorem alephIdx_aleph' (o : Ordinal) : (aleph' o).alephIdx = o := Cardinal.alephIdx.relIso.toEquiv.apply_symm_apply o @@ -608,7 +608,7 @@ theorem aleph'_limit {o : Ordinal} (ho : o.IsLimit) : aleph' o = ⨆ a : Iio o, theorem aleph'_omega0 : aleph' ω = ℵ₀ := preAleph_omega0 -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias aleph'_omega := aleph'_omega0 /-- `aleph'` and `aleph_idx` form an equivalence between `Ordinal` and `Cardinal` -/ @@ -638,7 +638,7 @@ theorem aleph'_isNormal : IsNormal (ord ∘ aleph') := -- They should also use `¬ BddAbove` instead of `Unbounded (· < ·)`. /-- Ordinals that are cardinals are unbounded. -/ -@[deprecated (since := "2024-09-24")] +@[deprecated "No deprecation message was provided." (since := "2024-09-24")] theorem ord_card_unbounded : Unbounded (· < ·) { b : Ordinal | b.card.ord = b } := unbounded_lt_iff.2 fun a => ⟨_, @@ -646,16 +646,16 @@ theorem ord_card_unbounded : Unbounded (· < ·) { b : Ordinal | b.card.ord = b dsimp rw [card_ord], (lt_ord_succ_card a).le⟩⟩ -@[deprecated (since := "2024-09-24")] +@[deprecated "No deprecation message was provided." (since := "2024-09-24")] theorem eq_aleph'_of_eq_card_ord {o : Ordinal} (ho : o.card.ord = o) : ∃ a, (aleph' a).ord = o := ⟨aleph'.symm o.card, by simpa using ho⟩ /-- Infinite ordinals that are cardinals are unbounded. -/ -@[deprecated (since := "2024-09-24")] +@[deprecated "No deprecation message was provided." (since := "2024-09-24")] theorem ord_card_unbounded' : Unbounded (· < ·) { b : Ordinal | b.card.ord = b ∧ ω ≤ b } := (unbounded_lt_inter_le ω).2 ord_card_unbounded -@[deprecated (since := "2024-09-24")] +@[deprecated "No deprecation message was provided." (since := "2024-09-24")] theorem eq_aleph_of_eq_card_ord {o : Ordinal} (ho : o.card.ord = o) (ho' : ω ≤ o) : ∃ a, (ℵ_ a).ord = o := by cases' eq_aleph'_of_eq_card_ord ho with a ha diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index 18f50f95db1cb..e0105b3587892 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -145,7 +145,7 @@ protected theorem eq : #α = #β ↔ Nonempty (α ≃ β) := Quotient.eq' /-- Avoid using `Quotient.mk` to construct a `Cardinal` directly -/ -@[deprecated (since := "2024-10-24")] +@[deprecated "No deprecation message was provided." (since := "2024-10-24")] theorem mk'_def (α : Type u) : @Eq Cardinal ⟦α⟧ #α := rfl diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index 87b3654599332..9a8363e186c9a 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -114,7 +114,7 @@ def StrictOrder.cof (r : α → α → Prop) : Cardinal := Order.cof (swap rᶜ) /-- The set in the definition of `Order.StrictOrder.cof` is nonempty. -/ -@[deprecated (since := "2024-10-22")] +@[deprecated "No deprecation message was provided." (since := "2024-10-22")] theorem StrictOrder.cof_nonempty (r : α → α → Prop) [IsIrrefl α r] : { c | ∃ S : Set α, Unbounded r S ∧ #S = c }.Nonempty := @Order.cof_nonempty α _ (IsRefl.swap rᶜ) diff --git a/Mathlib/SetTheory/Game/Ordinal.lean b/Mathlib/SetTheory/Game/Ordinal.lean index d281b72b36a02..cf5a7e8f4cf97 100644 --- a/Mathlib/SetTheory/Game/Ordinal.lean +++ b/Mathlib/SetTheory/Game/Ordinal.lean @@ -35,7 +35,7 @@ noncomputable def toPGame (o : Ordinal.{u}) : PGame.{u} := termination_by o decreasing_by exact ((enumIsoToType o).symm x).prop -@[deprecated (since := "2024-09-22")] +@[deprecated "No deprecation message was provided." (since := "2024-09-22")] theorem toPGame_def (o : Ordinal) : o.toPGame = ⟨o.toType, PEmpty, fun x => ((enumIsoToType o).symm x).val.toPGame, PEmpty.elim⟩ := by rw [toPGame] diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 6cc7dfe51df9e..3563908538d09 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -605,14 +605,14 @@ theorem one_add_omega0 : 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")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias one_add_omega := one_add_omega0 @[simp] 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")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias one_add_of_omega_le := one_add_of_omega0_le /-! ### Multiplication of ordinals -/ @@ -1185,7 +1185,7 @@ def sup {ι : Type u} (f : ι → Ordinal.{max u v}) : Ordinal.{max u v} := iSup f set_option linter.deprecated false in -@[deprecated (since := "2024-08-27")] +@[deprecated "No deprecation message was provided." (since := "2024-08-27")] theorem sSup_eq_sup {ι : Type u} (f : ι → Ordinal.{max u v}) : sSup (Set.range f) = sup.{_, v} f := rfl @@ -1252,14 +1252,15 @@ protected theorem lt_iSup_iff {ι} {f : ι → Ordinal.{u}} {a : Ordinal.{u}} [S a < iSup f ↔ ∃ i, a < f i := lt_ciSup_iff' (bddAbove_of_small _) -@[deprecated (since := "2024-11-12")] alias lt_iSup := lt_iSup_iff +@[deprecated "No deprecation message was provided." (since := "2024-11-12")] +alias lt_iSup := lt_iSup_iff set_option linter.deprecated false in @[deprecated Ordinal.lt_iSup (since := "2024-08-27")] theorem lt_sup {ι : Type u} {f : ι → Ordinal.{max u v}} {a} : a < sup.{_, v} f ↔ ∃ i, a < f i := by simpa only [not_forall, not_le] using not_congr (@sup_le_iff.{_, v} _ f a) -@[deprecated (since := "2024-08-27")] +@[deprecated "No deprecation message was provided." (since := "2024-08-27")] theorem ne_iSup_iff_lt_iSup {ι : Type u} {f : ι → Ordinal.{max u v}} : (∀ i, f i ≠ iSup f) ↔ ∀ i, f i < iSup f := forall_congr' fun i => (Ordinal.le_iSup f i).lt_iff_ne.symm @@ -1377,7 +1378,7 @@ theorem unbounded_range_of_sup_ge {α β : Type u} (r : α → α → Prop) [IsW unbounded_range_of_le_iSup r f h set_option linter.deprecated false in -@[deprecated (since := "2024-08-27")] +@[deprecated "No deprecation message was provided." (since := "2024-08-27")] theorem le_sup_shrink_equiv {s : Set Ordinal.{u}} (hs : Small.{u} s) (a) (ha : a ∈ s) : a ≤ sup.{u, u} fun x => ((@equivShrink s hs).symm x).val := by convert le_sup.{u, u} (fun x => ((@equivShrink s hs).symm x).val) ((@equivShrink s hs) ⟨a, ha⟩) @@ -1422,7 +1423,7 @@ theorem IsNormal.apply_of_isLimit {f : Ordinal.{u} → Ordinal.{v}} (H : IsNorma rw [← H.map_iSup, ho.iSup_Iio] set_option linter.deprecated false in -@[deprecated (since := "2024-08-27")] +@[deprecated "No deprecation message was provided." (since := "2024-08-27")] theorem sup_eq_sSup {s : Set Ordinal.{u}} (hs : Small.{u} s) : (sup.{u, u} fun x => (@equivShrink s hs).symm x) = sSup s := let hs' := bddAbove_iff_small.2 hs @@ -1986,12 +1987,12 @@ theorem IsNormal.eq_iff_zero_and_succ {f g : Ordinal.{u} → Ordinal.{u}} (hf : Deprecated. If you need this value explicitly, write it in terms of `iSup`. If you just want an upper bound for the image of `op`, use that `Iio a ×ˢ Iio b` is a small set. -/ -@[deprecated (since := "2024-10-11")] +@[deprecated "No deprecation message was provided." (since := "2024-10-11")] def blsub₂ (o₁ o₂ : Ordinal) (op : {a : Ordinal} → (a < o₁) → {b : Ordinal} → (b < o₂) → Ordinal) : Ordinal := lsub (fun x : o₁.toType × o₂.toType => op (typein_lt_self x.1) (typein_lt_self x.2)) -@[deprecated (since := "2024-10-11")] +@[deprecated "No deprecation message was provided." (since := "2024-10-11")] theorem lt_blsub₂ {o₁ o₂ : Ordinal} (op : {a : Ordinal} → (a < o₁) → {b : Ordinal} → (b < o₂) → Ordinal) {a b : Ordinal} (ha : a < o₁) (hb : b < o₂) : op ha hb < blsub₂ o₁ o₂ op := by @@ -2012,34 +2013,34 @@ set_option linter.deprecated false def mex {ι : Type u} (f : ι → Ordinal.{max u v}) : Ordinal := sInf (Set.range f)ᶜ -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem mex_not_mem_range {ι : Type u} (f : ι → Ordinal.{max u v}) : mex.{_, v} f ∉ Set.range f := csInf_mem (nonempty_compl_range.{_, v} f) -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem le_mex_of_forall {ι : Type u} {f : ι → Ordinal.{max u v}} {a : Ordinal} (H : ∀ b < a, ∃ i, f i = b) : a ≤ mex.{_, v} f := by by_contra! h exact mex_not_mem_range f (H _ h) -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem ne_mex {ι : Type u} (f : ι → Ordinal.{max u v}) : ∀ i, f i ≠ mex.{_, v} f := by simpa using mex_not_mem_range.{_, v} f -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem mex_le_of_ne {ι} {f : ι → Ordinal} {a} (ha : ∀ i, f i ≠ a) : mex f ≤ a := csInf_le' (by simp [ha]) -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem exists_of_lt_mex {ι} {f : ι → Ordinal} {a} (ha : a < mex f) : ∃ i, f i = a := by by_contra! ha' exact ha.not_le (mex_le_of_ne ha') -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem mex_le_lsub {ι : Type u} (f : ι → Ordinal.{max u v}) : mex.{_, v} f ≤ lsub.{_, v} f := csInf_le' (lsub_not_mem_range f) -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem mex_monotone {α β : Type u} {f : α → Ordinal.{max u v}} {g : β → Ordinal.{max u v}} (h : Set.range f ⊆ Set.range g) : mex.{_, v} f ≤ mex.{_, v} g := by refine mex_le_of_ne fun i hi => ?_ @@ -2072,18 +2073,18 @@ theorem mex_lt_ord_succ_mk {ι : Type u} (f : ι → Ordinal.{u}) : def bmex (o : Ordinal) (f : ∀ a < o, Ordinal) : Ordinal := mex (familyOfBFamily o f) -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem bmex_not_mem_brange {o : Ordinal} (f : ∀ a < o, Ordinal) : bmex o f ∉ brange o f := by rw [← range_familyOfBFamily] apply mex_not_mem_range -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem le_bmex_of_forall {o : Ordinal} (f : ∀ a < o, Ordinal) {a : Ordinal} (H : ∀ b < a, ∃ i hi, f i hi = b) : a ≤ bmex o f := by by_contra! h exact bmex_not_mem_brange f (H _ h) -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem ne_bmex {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) {i} (hi) : f i hi ≠ bmex.{_, v} o f := by convert (config := {transparency := .default}) @@ -2091,23 +2092,23 @@ theorem ne_bmex {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) {i} (hi) : -- Porting note: `familyOfBFamily_enum` → `typein_enum` rw [typein_enum] -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem bmex_le_of_ne {o : Ordinal} {f : ∀ a < o, Ordinal} {a} (ha : ∀ i hi, f i hi ≠ a) : bmex o f ≤ a := mex_le_of_ne fun _i => ha _ _ -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem exists_of_lt_bmex {o : Ordinal} {f : ∀ a < o, Ordinal} {a} (ha : a < bmex o f) : ∃ i hi, f i hi = a := by cases' exists_of_lt_mex ha with i hi exact ⟨_, typein_lt_self i, hi⟩ -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem bmex_le_blsub {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) : bmex.{_, v} o f ≤ blsub.{_, v} o f := mex_le_lsub _ -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem bmex_monotone {o o' : Ordinal.{u}} {f : ∀ a < o, Ordinal.{max u v}} {g : ∀ a < o', Ordinal.{max u v}} (h : brange o f ⊆ brange o' g) : bmex.{_, v} o f ≤ bmex.{_, v} o' g := @@ -2165,7 +2166,7 @@ theorem one_add_natCast (m : ℕ) : 1 + (m : Ordinal) = succ m := by rw [← Nat.cast_one, ← Nat.cast_add, add_comm] rfl -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias one_add_nat_cast := one_add_natCast -- See note [no_index around OfNat.ofNat] @@ -2179,43 +2180,43 @@ theorem natCast_mul (m : ℕ) : ∀ n : ℕ, ((m * n : ℕ) : Ordinal) = m * n | 0 => by simp | n + 1 => by rw [Nat.mul_succ, Nat.cast_add, natCast_mul m n, Nat.cast_succ, mul_add_one] -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_mul := natCast_mul @[deprecated Nat.cast_le (since := "2024-10-17")] theorem natCast_le {m n : ℕ} : (m : Ordinal) ≤ n ↔ m ≤ n := Nat.cast_le -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_le := natCast_le @[deprecated Nat.cast_inj (since := "2024-10-17")] theorem natCast_inj {m n : ℕ} : (m : Ordinal) = n ↔ m = n := Nat.cast_inj -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_inj := natCast_inj @[deprecated Nat.cast_lt (since := "2024-10-17")] theorem natCast_lt {m n : ℕ} : (m : Ordinal) < n ↔ m < n := Nat.cast_lt -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_lt := natCast_lt @[deprecated Nat.cast_eq_zero (since := "2024-10-17")] theorem natCast_eq_zero {n : ℕ} : (n : Ordinal) = 0 ↔ n = 0 := Nat.cast_eq_zero -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_eq_zero := natCast_eq_zero @[deprecated Nat.cast_ne_zero (since := "2024-10-17")] theorem natCast_ne_zero {n : ℕ} : (n : Ordinal) ≠ 0 ↔ n ≠ 0 := Nat.cast_ne_zero -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_ne_zero := natCast_ne_zero @[deprecated Nat.cast_pos' (since := "2024-10-17")] theorem natCast_pos {n : ℕ} : (0 : Ordinal) < n ↔ 0 < n := Nat.cast_pos' -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_pos := natCast_pos @[simp, norm_cast] @@ -2226,7 +2227,7 @@ theorem natCast_sub (m n : ℕ) : ((m - n : ℕ) : Ordinal) = m - n := by · apply (add_left_cancel n).1 rw [← Nat.cast_add, add_tsub_cancel_of_le h, Ordinal.add_sub_cancel_of_le (Nat.cast_le.2 h)] -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_sub := natCast_sub @[simp, norm_cast] @@ -2241,7 +2242,7 @@ theorem natCast_div (m n : ℕ) : ((m / n : ℕ) : Ordinal) = m / n := by ← Nat.div_lt_iff_lt_mul (Nat.pos_of_ne_zero hn)] apply Nat.lt_succ_self -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_div := natCast_div @[simp, norm_cast] @@ -2249,7 +2250,7 @@ theorem natCast_mod (m n : ℕ) : ((m % n : ℕ) : Ordinal) = m % n := by rw [← add_left_cancel, div_add_mod, ← natCast_div, ← natCast_mul, ← Nat.cast_add, Nat.div_add_mod] -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_mod := natCast_mod @[simp] @@ -2257,7 +2258,7 @@ theorem lift_natCast : ∀ n : ℕ, lift.{u, v} n = n | 0 => by simp | n + 1 => by simp [lift_natCast n] -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias lift_nat_cast := lift_natCast -- See note [no_index around OfNat.ofNat] @@ -2292,13 +2293,13 @@ theorem lt_add_of_limit {a b c : Ordinal.{u}} (h : IsLimit c) : theorem lt_omega0 {o : Ordinal} : o < ω ↔ ∃ n : ℕ, o = n := by simp_rw [← Cardinal.ord_aleph0, Cardinal.lt_ord, lt_aleph0, card_eq_nat] -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias lt_omega := lt_omega0 theorem nat_lt_omega0 (n : ℕ) : ↑n < ω := lt_omega0.2 ⟨_, rfl⟩ -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias nat_lt_omega := nat_lt_omega0 theorem omega0_pos : 0 < ω := @@ -2307,12 +2308,12 @@ theorem omega0_pos : 0 < ω := theorem omega0_ne_zero : ω ≠ 0 := omega0_pos.ne' -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias omega_ne_zero := omega0_ne_zero theorem one_lt_omega0 : 1 < ω := by simpa only [Nat.cast_one] using nat_lt_omega0 1 -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias one_lt_omega := one_lt_omega0 theorem isLimit_omega0 : IsLimit ω := @@ -2320,10 +2321,10 @@ theorem isLimit_omega0 : IsLimit ω := let ⟨n, e⟩ := lt_omega0.1 h rw [e]; exact nat_lt_omega0 (n + 1)⟩ -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] alias omega0_isLimit := isLimit_omega0 -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias omega_isLimit := isLimit_omega0 theorem omega0_le {o : Ordinal} : ω ≤ o ↔ ∀ n : ℕ, ↑n ≤ o := @@ -2332,7 +2333,7 @@ theorem omega0_le {o : Ordinal} : ω ≤ o ↔ ∀ n : ℕ, ↑n ≤ o := let ⟨n, e⟩ := lt_omega0.1 h rw [e, ← succ_le_iff]; exact H (n + 1)⟩ -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias omega_le := omega0_le @[simp] @@ -2344,7 +2345,7 @@ set_option linter.deprecated false in theorem sup_natCast : sup Nat.cast = ω := iSup_natCast -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias sup_nat_cast := sup_natCast theorem nat_lt_limit {o} (h : IsLimit o) : ∀ n : ℕ, ↑n < o @@ -2354,7 +2355,7 @@ theorem nat_lt_limit {o} (h : IsLimit o) : ∀ n : ℕ, ↑n < o theorem omega0_le_of_isLimit {o} (h : IsLimit o) : ω ≤ o := omega0_le.2 fun n => le_of_lt <| nat_lt_limit h n -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (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 @@ -2372,7 +2373,7 @@ theorem isLimit_iff_omega0_dvd {a : Ordinal} : IsLimit a ↔ a ≠ 0 ∧ ω ∣ intro e simp only [e, mul_zero] -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (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) @@ -2416,7 +2417,7 @@ theorem add_le_of_forall_add_lt {a b c : Ordinal} (hb : 0 < b) (h : ∀ d < b, a 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")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias IsNormal.apply_omega := IsNormal.apply_omega0 @[simp] @@ -2460,7 +2461,7 @@ theorem isLimit_ord {c} (co : ℵ₀ ≤ c) : (ord c).IsLimit := by · rw [ord_aleph0] exact Ordinal.isLimit_omega0 -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] alias ord_isLimit := isLimit_ord theorem noMaxOrder {c} (h : ℵ₀ ≤ c) : NoMaxOrder c.ord.toType := diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index 5a31bfe919ee9..ec9a34419fbf4 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -387,7 +387,7 @@ def typein (r : α → α → Prop) [IsWellOrder α r] : @PrincipalSeg α Ordina alias typein.principalSeg := typein set_option linter.deprecated false in -@[deprecated (since := "2024-10-09")] +@[deprecated "No deprecation message was provided." (since := "2024-10-09")] theorem typein.principalSeg_coe (r : α → α → Prop) [IsWellOrder α r] : (typein.principalSeg r : α → Ordinal) = typein r := rfl @@ -513,7 +513,7 @@ noncomputable def enumIsoToType (o : Ordinal) : Set.Iio o ≃o o.toType where right_inv _ := enum_typein _ _ map_rel_iff' := enum_le_enum' _ -@[deprecated (since := "2024-08-26")] +@[deprecated "No deprecation message was provided." (since := "2024-08-26")] alias enumIsoOut := enumIsoToType instance small_Iio (o : Ordinal.{u}) : Small.{u} (Iio o) := @@ -925,7 +925,7 @@ theorem card_succ (o : Ordinal) : card (succ o) = card o + 1 := by theorem natCast_succ (n : ℕ) : ↑n.succ = succ (n : Ordinal) := rfl -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_succ := natCast_succ instance uniqueIioOne : Unique (Iio (1 : Ordinal)) where @@ -1213,7 +1213,7 @@ alias card_typein_out_lt := card_typein_toType_lt theorem mk_Iio_ord_toType {c : Cardinal} (i : c.ord.toType) : #(Iio i) < c := card_typein_toType_lt c i -@[deprecated (since := "2024-08-26")] +@[deprecated "No deprecation message was provided." (since := "2024-08-26")] alias mk_Iio_ord_out_α := mk_Iio_ord_toType theorem ord_injective : Injective ord := by diff --git a/Mathlib/SetTheory/Ordinal/Enum.lean b/Mathlib/SetTheory/Ordinal/Enum.lean index 4e5d4d38a7cf6..e8b3bb9314bc5 100644 --- a/Mathlib/SetTheory/Ordinal/Enum.lean +++ b/Mathlib/SetTheory/Ordinal/Enum.lean @@ -33,7 +33,7 @@ termination_by o variable {s : Set Ordinal.{u}} -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem enumOrd_def (o : Ordinal.{u}) : enumOrd s o = sInf (s ∩ { b | ∀ c, c < o → enumOrd s c < b }) := by rw [enumOrd] diff --git a/Mathlib/SetTheory/Ordinal/FixedPoint.lean b/Mathlib/SetTheory/Ordinal/FixedPoint.lean index 36b50a052b83d..84f2f8db36b0d 100644 --- a/Mathlib/SetTheory/Ordinal/FixedPoint.lean +++ b/Mathlib/SetTheory/Ordinal/FixedPoint.lean @@ -52,7 +52,7 @@ least `a`, and `Ordinal.nfpFamily_le_fp` shows this is the least ordinal with th def nfpFamily (f : ι → Ordinal.{u} → Ordinal.{u}) (a : Ordinal.{u}) : Ordinal := ⨆ i, List.foldr f a i -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpFamily_eq_sup (f : ι → Ordinal.{u} → Ordinal.{u}) (a : Ordinal.{u}) : nfpFamily f a = ⨆ i, List.foldr f a i := rfl @@ -238,48 +238,48 @@ def nfpBFamily (o : Ordinal.{u}) (f : ∀ b < o, Ordinal.{max u v} → Ordinal.{ Ordinal.{max u v} → Ordinal.{max u v} := nfpFamily (familyOfBFamily o f) -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_eq_nfpFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) : nfpBFamily.{u, v} o f = nfpFamily (familyOfBFamily o f) := rfl -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem foldr_le_nfpBFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) (a l) : List.foldr (familyOfBFamily o f) a l ≤ nfpBFamily.{u, v} o f a := Ordinal.le_iSup _ _ -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem le_nfpBFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) (a) : a ≤ nfpBFamily.{u, v} o f a := Ordinal.le_iSup (fun _ ↦ List.foldr _ a _) [] -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem lt_nfpBFamily {a b} : a < nfpBFamily.{u, v} o f b ↔ ∃ l, a < List.foldr (familyOfBFamily o f) b l := Ordinal.lt_iSup_iff -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_le_iff {o : Ordinal} {f : ∀ b < o, Ordinal → Ordinal} {a b} : nfpBFamily.{u, v} o f a ≤ b ↔ ∀ l, List.foldr (familyOfBFamily o f) a l ≤ b := Ordinal.iSup_le_iff -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_le {o : Ordinal} {f : ∀ b < o, Ordinal → Ordinal} {a b} : (∀ l, List.foldr (familyOfBFamily o f) a l ≤ b) → nfpBFamily.{u, v} o f a ≤ b := Ordinal.iSup_le -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_monotone (hf : ∀ i hi, Monotone (f i hi)) : Monotone (nfpBFamily.{u, v} o f) := nfpFamily_monotone fun _ => hf _ _ -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem apply_lt_nfpBFamily (H : ∀ i hi, IsNormal (f i hi)) {a b} (hb : b < nfpBFamily.{u, v} o f a) (i hi) : f i hi b < nfpBFamily.{u, v} o f a := by rw [← familyOfBFamily_enum o f] apply apply_lt_nfpFamily (fun _ => H _ _) hb -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem apply_lt_nfpBFamily_iff (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a b} : (∀ i hi, f i hi b < nfpBFamily.{u, v} o f a) ↔ b < nfpBFamily.{u, v} o f a := ⟨fun h => by @@ -287,19 +287,19 @@ theorem apply_lt_nfpBFamily_iff (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) refine (apply_lt_nfpFamily_iff ?_).1 fun _ => h _ _ exact fun _ => H _ _, apply_lt_nfpBFamily H⟩ -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_le_apply (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a b} : (∃ i hi, nfpBFamily.{u, v} o f a ≤ f i hi b) ↔ nfpBFamily.{u, v} o f a ≤ b := by rw [← not_iff_not] push_neg exact apply_lt_nfpBFamily_iff.{u, v} ho H -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_le_fp (H : ∀ i hi, Monotone (f i hi)) {a b} (ab : a ≤ b) (h : ∀ i hi, f i hi b ≤ b) : nfpBFamily.{u, v} o f a ≤ b := nfpFamily_le_fp (fun _ => H _ _) ab fun _ => h _ _ -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_fp {i hi} (H : IsNormal (f i hi)) (a) : f i hi (nfpBFamily.{u, v} o f a) = nfpBFamily.{u, v} o f a := by rw [← familyOfBFamily_enum o f] @@ -307,7 +307,7 @@ theorem nfpBFamily_fp {i hi} (H : IsNormal (f i hi)) (a) : rw [familyOfBFamily_enum] exact H -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem apply_le_nfpBFamily (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a b} : (∀ i hi, f i hi b ≤ nfpBFamily.{u, v} o f a) ↔ b ≤ nfpBFamily.{u, v} o f a := by refine ⟨fun h => ?_, fun h i hi => ?_⟩ @@ -316,13 +316,13 @@ theorem apply_le_nfpBFamily (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a · rw [← nfpBFamily_fp (H i hi)] exact (H i hi).monotone h -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_eq_self {a} (h : ∀ i hi, f i hi a = a) : nfpBFamily.{u, v} o f a = a := nfpFamily_eq_self fun _ => h _ _ /-- A generalization of the fixed point lemma for normal functions: any family of normal functions has an unbounded set of common fixed points. -/ -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem not_bddAbove_fp_bfamily (H : ∀ i hi, IsNormal (f i hi)) : ¬ BddAbove (⋂ (i) (hi), Function.fixedPoints (f i hi)) := by rw [not_bddAbove_iff] @@ -347,12 +347,12 @@ def derivBFamily (o : Ordinal.{u}) (f : ∀ b < o, Ordinal.{max u v} → Ordinal Ordinal.{max u v} → Ordinal.{max u v} := derivFamily (familyOfBFamily o f) -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem derivBFamily_eq_derivFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) : derivBFamily.{u, v} o f = derivFamily (familyOfBFamily o f) := rfl -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem isNormal_derivBFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) : IsNormal (derivBFamily o f) := isNormal_derivFamily _ @@ -360,7 +360,7 @@ theorem isNormal_derivBFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) @[deprecated isNormal_derivBFamily (since := "2024-10-11")] alias derivBFamily_isNormal := isNormal_derivBFamily -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem derivBFamily_fp {i hi} (H : IsNormal (f i hi)) (a : Ordinal) : f i hi (derivBFamily.{u, v} o f a) = derivBFamily.{u, v} o f a := by rw [← familyOfBFamily_enum o f] @@ -368,7 +368,7 @@ theorem derivBFamily_fp {i hi} (H : IsNormal (f i hi)) (a : Ordinal) : rw [familyOfBFamily_enum] exact H -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem le_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : (∀ i hi, f i hi a ≤ a) ↔ ∃ b, derivBFamily.{u, v} o f b = a := by unfold derivBFamily @@ -378,7 +378,7 @@ theorem le_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : apply h · exact fun _ => H _ _ -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem fp_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : (∀ i hi, f i hi a = a) ↔ ∃ b, derivBFamily.{u, v} o f b = a := by rw [← le_iff_derivBFamily H] @@ -387,7 +387,7 @@ theorem fp_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : exact h i hi /-- For a family of normal functions, `Ordinal.derivBFamily` enumerates the common fixed points. -/ -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem derivBFamily_eq_enumOrd (H : ∀ i hi, IsNormal (f i hi)) : derivBFamily.{u, v} o f = enumOrd (⋂ (i) (hi), Function.fixedPoints (f i hi)) := by rw [eq_comm, eq_enumOrd _ (not_bddAbove_fp_bfamily H)] @@ -428,7 +428,7 @@ theorem iSup_iterate_eq_nfp (f : Ordinal.{u} → Ordinal.{u}) (a : Ordinal.{u}) exact Ordinal.le_iSup _ _ set_option linter.deprecated false in -@[deprecated (since := "2024-08-27")] +@[deprecated "No deprecation message was provided." (since := "2024-08-27")] theorem sup_iterate_eq_nfp (f : Ordinal.{u} → Ordinal.{u}) (a : Ordinal.{u}) : (sup fun n : ℕ => f^[n] a) = nfp f a := by refine le_antisymm ?_ (sup_le fun l => ?_) @@ -580,7 +580,7 @@ theorem nfp_add_eq_mul_omega0 {a b} (hba : b ≤ a * ω) : nfp (a + ·) b = a * exact nfp_monotone (isNormal_add_right a).monotone (Ordinal.zero_le b) · dsimp; rw [← mul_one_add, one_add_omega0] -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (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 @@ -593,14 +593,14 @@ theorem add_eq_right_iff_mul_omega0_le {a b : Ordinal} : a + b = b ↔ a * ω nth_rw 1 [← this] rwa [← add_assoc, ← mul_one_add, one_add_omega0] -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (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_omega0_le {a b : Ordinal} : a + b ≤ b ↔ a * ω ≤ b := by rw [← add_eq_right_iff_mul_omega0_le] exact (isNormal_add_right a).le_iff_eq -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (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 @@ -612,7 +612,7 @@ theorem deriv_add_eq_mul_omega0_add (a b : Ordinal.{u}) : deriv (a + ·) b = a * · rw [deriv_succ, h, add_succ] exact nfp_eq_self (add_eq_right_iff_mul_omega0_le.2 ((le_add_right _ _).trans (le_succ _))) -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias deriv_add_eq_mul_omega_add := deriv_add_eq_mul_omega0_add /-! ### Fixed points of multiplication -/ @@ -645,7 +645,7 @@ theorem nfp_mul_eq_opow_omega0 {a b : Ordinal} (hb : 0 < b) (hba : b ≤ a ^ ω) rw [← nfp_mul_one ha] exact nfp_monotone (isNormal_mul_right ha).monotone (one_le_iff_pos.2 hb) -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (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) : @@ -659,7 +659,7 @@ theorem eq_zero_or_opow_omega0_le_of_mul_eq_right {a b : Ordinal} (hab : a * b = rw [← Ne, ← one_le_iff_ne_zero] at hb exact nfp_le_fp (isNormal_mul_right ha).monotone hb (le_of_eq hab) -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (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 @@ -677,7 +677,7 @@ theorem mul_eq_right_iff_opow_omega0_dvd {a b : Ordinal} : a * b = b ↔ a ^ ω cases' h with c hc rw [hc, ← mul_assoc, ← opow_one_add, one_add_omega0] -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (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_omega0_dvd {a b : Ordinal} (ha : 0 < a) : @@ -685,7 +685,7 @@ theorem mul_le_right_iff_opow_omega0_dvd {a b : Ordinal} (ha : 0 < a) : rw [← mul_eq_right_iff_opow_omega0_dvd] exact (isNormal_mul_right ha).le_iff_eq -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (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) @@ -705,7 +705,7 @@ theorem nfp_mul_opow_omega0_add {a c : Ordinal} (b) (ha : 0 < a) (hc : 0 < c) rw [add_zero, mul_lt_mul_iff_left (opow_pos ω ha)] at this rwa [succ_le_iff] -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (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) : @@ -718,7 +718,7 @@ theorem deriv_mul_eq_opow_omega0_mul {a : Ordinal.{u}} (ha : 0 < a) (b) : · rw [deriv_succ, h] exact nfp_mul_opow_omega0_add c ha zero_lt_one (one_le_iff_pos.2 (opow_pos _ ha)) -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (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/Principal.lean b/Mathlib/SetTheory/Ordinal/Principal.lean index 5d6417f174042..194abeabc7c50 100644 --- a/Mathlib/SetTheory/Ordinal/Principal.lean +++ b/Mathlib/SetTheory/Ordinal/Principal.lean @@ -129,7 +129,7 @@ theorem not_bddAbove_principal (op : Ordinal → Ordinal → Ordinal) : exact ((le_nfp _ _).trans (ha (principal_nfp_iSup op (succ a)))).not_lt (lt_succ a) set_option linter.deprecated false in -@[deprecated (since := "2024-10-11")] +@[deprecated "No deprecation message was provided." (since := "2024-10-11")] theorem principal_nfp_blsub₂ (op : Ordinal → Ordinal → Ordinal) (o : Ordinal) : Principal op (nfp (fun o' => blsub₂.{u, u, u} o' o' (@fun a _ b _ => op a b)) o) := by intro a b ha hb @@ -147,7 +147,7 @@ theorem principal_nfp_blsub₂ (op : Ordinal → Ordinal → Ordinal) (o : Ordin exact lt_blsub₂ (@fun a _ b _ => op a b) hm (hn.trans_le h) set_option linter.deprecated false in -@[deprecated (since := "2024-10-11")] +@[deprecated "No deprecation message was provided." (since := "2024-10-11")] theorem unbounded_principal (op : Ordinal → Ordinal → Ordinal) : Set.Unbounded (· < ·) { o | Principal op o } := fun o => ⟨_, principal_nfp_blsub₂ op o, (le_nfp _ o).not_lt⟩ diff --git a/Mathlib/SetTheory/ZFC/Basic.lean b/Mathlib/SetTheory/ZFC/Basic.lean index c7a0c35cdfcef..a2b3354096bf0 100644 --- a/Mathlib/SetTheory/ZFC/Basic.lean +++ b/Mathlib/SetTheory/ZFC/Basic.lean @@ -447,12 +447,12 @@ set_option linter.deprecated false /-- Function equivalence is defined so that `f ~ g` iff `∀ x y, x ~ y → f x ~ g y`. This extends to equivalence of `n`-ary functions. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Arity.Equiv : ∀ {n}, OfArity PSet.{u} PSet.{u} n → OfArity PSet.{u} PSet.{u} n → Prop | 0, a, b => PSet.Equiv a b | _ + 1, a, b => ∀ x y : PSet, PSet.Equiv x y → Arity.Equiv (a x) (b y) -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] theorem Arity.equiv_const {a : PSet.{u}} : ∀ n, Arity.Equiv (OfArity.const PSet.{u} a n) (OfArity.const PSet.{u} a n) | 0 => Equiv.rfl @@ -460,46 +460,46 @@ theorem Arity.equiv_const {a : PSet.{u}} : /-- `resp n` is the collection of n-ary functions on `PSet` that respect equivalence, i.e. when the inputs are equivalent the output is as well. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Resp (n) := { x : OfArity PSet.{u} PSet.{u} n // Arity.Equiv x x } -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] instance Resp.inhabited {n} : Inhabited (Resp n) := ⟨⟨OfArity.const _ default _, Arity.equiv_const _⟩⟩ /-- The `n`-ary image of a `(n + 1)`-ary function respecting equivalence as a function respecting equivalence. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Resp.f {n} (f : Resp (n + 1)) (x : PSet) : Resp n := ⟨f.1 x, f.2 _ _ <| Equiv.refl x⟩ /-- Function equivalence for functions respecting equivalence. See `PSet.Arity.Equiv`. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Resp.Equiv {n} (a b : Resp n) : Prop := Arity.Equiv a.1 b.1 -@[deprecated (since := "2024-09-02"), refl] +@[deprecated "No deprecation message was provided." (since := "2024-09-02"), refl] protected theorem Resp.Equiv.refl {n} (a : Resp n) : Resp.Equiv a a := a.2 -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] protected theorem Resp.Equiv.euc : ∀ {n} {a b c : Resp n}, Resp.Equiv a b → Resp.Equiv c b → Resp.Equiv a c | 0, _, _, _, hab, hcb => PSet.Equiv.euc hab hcb | n + 1, a, b, c, hab, hcb => fun x y h => @Resp.Equiv.euc n (a.f x) (b.f y) (c.f y) (hab _ _ h) (hcb _ _ <| PSet.Equiv.refl y) -@[deprecated (since := "2024-09-02"), symm] +@[deprecated "No deprecation message was provided." (since := "2024-09-02"), symm] protected theorem Resp.Equiv.symm {n} {a b : Resp n} : Resp.Equiv a b → Resp.Equiv b a := (Resp.Equiv.refl b).euc -@[deprecated (since := "2024-09-02"), trans] +@[deprecated "No deprecation message was provided." (since := "2024-09-02"), trans] protected theorem Resp.Equiv.trans {n} {x y z : Resp n} (h1 : Resp.Equiv x y) (h2 : Resp.Equiv y z) : Resp.Equiv x z := h1.euc h2.symm -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] instance Resp.setoid {n} : Setoid (Resp n) := ⟨Resp.Equiv, Resp.Equiv.refl, Resp.Equiv.symm, Resp.Equiv.trans⟩ @@ -611,7 +611,7 @@ set_option linter.deprecated false namespace Resp /-- Helper function for `PSet.eval`. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] def evalAux : ∀ {n}, { f : Resp n → OfArity ZFSet.{u} ZFSet.{u} n // ∀ a b : Resp n, Resp.Equiv a b → f a = f b } @@ -626,11 +626,11 @@ def evalAux : evalAux.2 (Resp.f b z) (Resp.f c z) (h _ _ (PSet.Equiv.refl z))⟩ /-- An equivalence-respecting function yields an n-ary ZFC set function. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] def eval (n) : Resp n → OfArity ZFSet.{u} ZFSet.{u} n := evalAux.1 -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] theorem eval_val {n f x} : (@eval (n + 1) f : ZFSet → OfArity ZFSet ZFSet n) ⟦x⟧ = eval n (Resp.f f x) := rfl @@ -640,24 +640,25 @@ end Resp /-- A set function is "definable" if it is the image of some n-ary pre-set function. This isn't exactly definability, but is useful as a sufficient condition for functions that have a computable image. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] class inductive Definable (n) : OfArity ZFSet.{u} ZFSet.{u} n → Type (u + 1) | mk (f) : Definable n (Resp.eval n f) -attribute [deprecated (since := "2024-09-02"), instance] Definable.mk +attribute [deprecated "No deprecation message was provided." (since := "2024-09-02"), instance] + Definable.mk /-- The evaluation of a function respecting equivalence is definable, by that same function. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Definable.EqMk {n} (f) : ∀ {s : OfArity ZFSet.{u} ZFSet.{u} n} (_ : Resp.eval _ f = s), Definable n s | _, rfl => ⟨f⟩ /-- Turns a definable function into a function that respects equivalence. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Definable.Resp {n} : ∀ (s : OfArity ZFSet.{u} ZFSet.{u} n) [Definable n s], Resp n | _, ⟨f⟩ => f -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] theorem Definable.eq {n} : ∀ (s : OfArity ZFSet.{u} ZFSet.{u} n) [H : Definable n s], (@Definable.Resp n s H).eval _ = s | _, ⟨_⟩ => rfl @@ -670,7 +671,7 @@ open PSet ZFSet set_option linter.deprecated false in /-- All functions are classically definable. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] noncomputable def allDefinable : ∀ {n} (F : OfArity ZFSet ZFSet n), Definable n F | 0, F => let p := @Quotient.exists_rep PSet _ F @@ -706,7 +707,7 @@ theorem exact {x y : PSet} : mk x = mk y → PSet.Equiv x y := Quotient.exact set_option linter.deprecated false in -@[deprecated (since := "2024-09-02"), simp] +@[deprecated "No deprecation message was provided." (since := "2024-09-02"), simp] theorem eval_mk {n f x} : (@Resp.eval (n + 1) f : ZFSet → OfArity ZFSet ZFSet n) (mk x) = Resp.eval n (Resp.f f x) := rfl diff --git a/Mathlib/Topology/Algebra/ConstMulAction.lean b/Mathlib/Topology/Algebra/ConstMulAction.lean index 1da40dadc219d..44b6433923f59 100644 --- a/Mathlib/Topology/Algebra/ConstMulAction.lean +++ b/Mathlib/Topology/Algebra/ConstMulAction.lean @@ -266,7 +266,7 @@ theorem smul_mem_nhds_smul_iff {t : Set α} (g : G) {a : α} : g • t ∈ 𝓝 @[to_additive] alias ⟨_, smul_mem_nhds_smul⟩ := smul_mem_nhds_smul_iff -@[to_additive (attr := deprecated (since := "2024-08-06"))] +@[to_additive (attr := deprecated "No deprecation message was provided." (since := "2024-08-06"))] alias smul_mem_nhds := smul_mem_nhds_smul @[to_additive (attr := simp)] diff --git a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean index 91b1a4ca3bcd4..f7373a1da6da0 100644 --- a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean +++ b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean @@ -65,7 +65,8 @@ homomorphisms. Deprecated and changed from a `class` to a `structure`. Use `[MonoidHomClass F A B] [ContinuousMapClass F A B]` instead. -/ -@[to_additive (attr := deprecated (since := "2024-10-08"))] +@[to_additive (attr := deprecated "Use `[MonoidHomClass F A B] [ContinuousMapClass F A B]` instead." + (since := "2024-10-08"))] structure ContinuousMonoidHomClass (A B : outParam Type*) [Monoid A] [Monoid B] [TopologicalSpace A] [TopologicalSpace B] [FunLike F A B] extends MonoidHomClass F A B, ContinuousMapClass F A B : Prop diff --git a/Mathlib/Topology/Algebra/Group/Quotient.lean b/Mathlib/Topology/Algebra/Group/Quotient.lean index 4a2b490786409..e46af523bf2ca 100644 --- a/Mathlib/Topology/Algebra/Group/Quotient.lean +++ b/Mathlib/Topology/Algebra/Group/Quotient.lean @@ -76,7 +76,7 @@ instance instLocallyCompactSpace [LocallyCompactSpace G] (N : Subgroup G) : LocallyCompactSpace (G ⧸ N) := QuotientGroup.isOpenQuotientMap_mk.locallyCompactSpace -@[to_additive (attr := deprecated (since := "2024-10-05"))] +@[to_additive (attr := deprecated "No deprecation message was provided." (since := "2024-10-05"))] theorem continuous_smul₁ (x : G ⧸ N) : Continuous fun g : G => g • x := continuous_id.smul continuous_const @@ -101,7 +101,7 @@ instance instSecondCountableTopology [SecondCountableTopology G] : SecondCountableTopology (G ⧸ N) := ContinuousConstSMul.secondCountableTopology -@[to_additive (attr := deprecated (since := "2024-08-05"))] +@[to_additive (attr := deprecated "No deprecation message was provided." (since := "2024-08-05"))] theorem nhds_one_isCountablyGenerated [FirstCountableTopology G] [N.Normal] : (𝓝 (1 : G ⧸ N)).IsCountablyGenerated := inferInstance @@ -117,7 +117,7 @@ instance instTopologicalGroup [N.Normal] : TopologicalGroup (G ⧸ N) where exact continuous_mk.comp continuous_mul continuous_inv := continuous_inv.quotient_map' _ -@[to_additive (attr := deprecated (since := "2024-08-05"))] +@[to_additive (attr := deprecated "No deprecation message was provided." (since := "2024-08-05"))] theorem _root_.topologicalGroup_quotient [N.Normal] : TopologicalGroup (G ⧸ N) := instTopologicalGroup N diff --git a/Mathlib/Topology/LocalAtTarget.lean b/Mathlib/Topology/LocalAtTarget.lean index 9ab72631191fb..af16705f2ce51 100644 --- a/Mathlib/Topology/LocalAtTarget.lean +++ b/Mathlib/Topology/LocalAtTarget.lean @@ -82,7 +82,7 @@ theorem IsClosedMap.restrictPreimage (H : IsClosedMap f) (s : Set β) : simpa [isClosed_induced_iff] exact fun u hu e => ⟨f '' u, H u hu, by simp [← e, image_restrictPreimage]⟩ -@[deprecated (since := "2024-04-02")] +@[deprecated "No deprecation message was provided." (since := "2024-04-02")] theorem Set.restrictPreimage_isClosedMap (s : Set β) (H : IsClosedMap f) : IsClosedMap (s.restrictPreimage f) := H.restrictPreimage s @@ -94,7 +94,7 @@ theorem IsOpenMap.restrictPreimage (H : IsOpenMap f) (s : Set β) : simpa [isOpen_induced_iff] exact fun u hu e => ⟨f '' u, H u hu, by simp [← e, image_restrictPreimage]⟩ -@[deprecated (since := "2024-04-02")] +@[deprecated "No deprecation message was provided." (since := "2024-04-02")] theorem Set.restrictPreimage_isOpenMap (s : Set β) (H : IsOpenMap f) : IsOpenMap (s.restrictPreimage f) := H.restrictPreimage s diff --git a/Mathlib/Topology/MetricSpace/Polish.lean b/Mathlib/Topology/MetricSpace/Polish.lean index d1e851e3c7d44..93c053adc16d7 100644 --- a/Mathlib/Topology/MetricSpace/Polish.lean +++ b/Mathlib/Topology/MetricSpace/Polish.lean @@ -102,7 +102,7 @@ instance (priority := 100) instMetrizableSpace (α : Type*) [TopologicalSpace α letI := upgradePolishSpace α infer_instance -@[deprecated (since := "2024-02-23")] +@[deprecated "No deprecation message was provided." (since := "2024-02-23")] theorem t2Space (α : Type*) [TopologicalSpace α] [PolishSpace α] : T2Space α := inferInstance /-- A countable product of Polish spaces is Polish. -/ diff --git a/Mathlib/Topology/NoetherianSpace.lean b/Mathlib/Topology/NoetherianSpace.lean index 64689fd04f3f0..931e8748cbfce 100644 --- a/Mathlib/Topology/NoetherianSpace.lean +++ b/Mathlib/Topology/NoetherianSpace.lean @@ -98,7 +98,7 @@ theorem noetherianSpace_iff_isCompact : NoetherianSpace α ↔ ∀ s : Set α, I instance [NoetherianSpace α] : WellFoundedLT (Closeds α) := Iff.mp ((noetherianSpace_TFAE α).out 0 1) ‹_› -@[deprecated (since := "2024-10-07")] +@[deprecated "No deprecation message was provided." (since := "2024-10-07")] theorem NoetherianSpace.wellFounded_closeds [NoetherianSpace α] : WellFounded fun s t : Closeds α => s < t := wellFounded_lt diff --git a/Mathlib/Topology/Order/OrderClosed.lean b/Mathlib/Topology/Order/OrderClosed.lean index 7ded140f87a07..46e8626f54295 100644 --- a/Mathlib/Topology/Order/OrderClosed.lean +++ b/Mathlib/Topology/Order/OrderClosed.lean @@ -368,7 +368,7 @@ variable [TopologicalSpace α] [Preorder α] [ClosedIciTopology α] {f : β → theorem isClosed_Ici {a : α} : IsClosed (Ici a) := ClosedIciTopology.isClosed_Ici a -@[deprecated (since := "2024-02-15")] +@[deprecated "No deprecation message was provided." (since := "2024-02-15")] lemma ClosedIciTopology.isClosed_ge' (a : α) : IsClosed {x | a ≤ x} := isClosed_Ici a export ClosedIciTopology (isClosed_ge') From 7113817ab894ff56fb3a99f2efdff55f8da6df16 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 25 Nov 2024 13:34:39 +0000 Subject: [PATCH 289/829] chore: change the definition of `List.finRange` (#19447) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit François Dorais has been [working](https://github.com/leanprover-community/batteries/pull/1055) on upstreaming `List.finRange`, but wants to change the definition at the same time. That was running into difficulties, which hopefully are resolved here. --- Archive/Imo/Imo2024Q5.lean | 2 +- .../Complex/UpperHalfPlane/Basic.lean | 12 ++-- Mathlib/Analysis/Convex/StoneSeparation.lean | 4 +- .../Analysis/InnerProductSpace/TwoDim.lean | 33 +++++----- Mathlib/Data/List/Defs.lean | 12 +++- Mathlib/Data/List/FinRange.lean | 64 +++++++++++++++++-- Mathlib/Data/List/Range.lean | 49 -------------- Mathlib/Data/List/Sublists.lean | 3 +- .../AffineSpace/Independent.lean | 2 +- Mathlib/NumberTheory/Modular.lean | 18 +++--- .../NumberField/CanonicalEmbedding/Basic.lean | 20 +++--- Mathlib/Order/RelSeries.lean | 4 +- Mathlib/RingTheory/Complex.lean | 8 +-- Mathlib/SetTheory/Cardinal/Subfield.lean | 2 +- 14 files changed, 127 insertions(+), 106 deletions(-) diff --git a/Archive/Imo/Imo2024Q5.lean b/Archive/Imo/Imo2024Q5.lean index 7b090f48cc987..5cbdf23de02f8 100644 --- a/Archive/Imo/Imo2024Q5.lean +++ b/Archive/Imo/Imo2024Q5.lean @@ -509,7 +509,7 @@ lemma Strategy.play_two (s : Strategy N) (m : MonsterData N) {k : ℕ} (hk : 2 < fin_cases i · rfl · have h : (1 : Fin 2) = Fin.last 1 := rfl - simp only [Fin.snoc_zero, Nat.reduceAdd, Fin.mk_one, Fin.isValue, Matrix.cons_val_one, + simp only [Fin.snoc_zero, Nat.reduceAdd, Fin.mk_one, Fin.isValue, id_eq, Matrix.cons_val_one, Matrix.head_cons] simp only [h, Fin.snoc_last] convert rfl diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean index 0ac5af563ff8a..6a3125c87d0c9 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean @@ -243,9 +243,9 @@ theorem denom_cocycle (x y : GL(2, ℝ)⁺) (z : ℍ) : denom (x * y) z = denom x (smulAux y z) * denom y z := by change _ = (_ * (_ / _) + _) * _ field_simp [denom_ne_zero] - simp only [Matrix.mul_apply, dotProduct, Fin.sum_univ_succ, denom, num, Subgroup.coe_mul, - GeneralLinearGroup.coe_mul, Fintype.univ_ofSubsingleton, Fin.mk_zero, Finset.sum_singleton, - Fin.succ_zero_eq_one, Complex.ofReal_add, Complex.ofReal_mul] + simp only [denom, Subgroup.coe_mul, Fin.isValue, Units.val_mul, mul_apply, Fin.sum_univ_succ, + Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton, Fin.succ_zero_eq_one, + Complex.ofReal_add, Complex.ofReal_mul, num] ring theorem mul_smul' (x y : GL(2, ℝ)⁺) (z : ℍ) : smulAux (x * y) z = smulAux x (smulAux y z) := by @@ -254,9 +254,9 @@ theorem mul_smul' (x y : GL(2, ℝ)⁺) (z : ℍ) : smulAux (x * y) z = smulAux change _ / _ = (_ * (_ / _) + _) / _ rw [denom_cocycle] field_simp [denom_ne_zero] - simp only [Matrix.mul_apply, dotProduct, Fin.sum_univ_succ, num, denom, Subgroup.coe_mul, - GeneralLinearGroup.coe_mul, Fintype.univ_ofSubsingleton, Fin.mk_zero, Finset.sum_singleton, - Fin.succ_zero_eq_one, Complex.ofReal_add, Complex.ofReal_mul] + simp only [num, Subgroup.coe_mul, Fin.isValue, Units.val_mul, mul_apply, Fin.sum_univ_succ, + Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton, Fin.succ_zero_eq_one, + Complex.ofReal_add, Complex.ofReal_mul, denom] ring /-- The action of `GLPos 2 ℝ` on the upper half-plane by fractional linear transformations. -/ diff --git a/Mathlib/Analysis/Convex/StoneSeparation.lean b/Mathlib/Analysis/Convex/StoneSeparation.lean index 76a95705677f6..5a52d98fda2ea 100644 --- a/Mathlib/Analysis/Convex/StoneSeparation.lean +++ b/Mathlib/Analysis/Convex/StoneSeparation.lean @@ -67,8 +67,8 @@ theorem not_disjoint_segment_convexHull_triple {p q u v x y z : E} (hz : z ∈ s ((az * av * bu) • p + ((bz * au * bv) • q + (au * av) • (az • x + bz • y))) · module congr 3 - simp only [w, z, smul_add, List.foldr, Matrix.cons_val_succ', Fin.mk_one, - Matrix.cons_val_one, Matrix.head_cons, add_zero] + simp only [smul_add, List.foldr, Fin.reduceFinMk, id_eq, Fin.isValue, Matrix.cons_val_two, + Nat.succ_eq_add_one, Nat.reduceAdd, Matrix.tail_cons, Matrix.head_cons, add_zero, w, z] /-- **Stone's Separation Theorem** -/ theorem exists_convex_convex_compl_subset (hs : Convex 𝕜 s) (ht : Convex 𝕜 t) (hst : Disjoint s t) : diff --git a/Mathlib/Analysis/InnerProductSpace/TwoDim.lean b/Mathlib/Analysis/InnerProductSpace/TwoDim.lean index b07b1d09381f7..afdbc274d603b 100644 --- a/Mathlib/Analysis/InnerProductSpace/TwoDim.lean +++ b/Mathlib/Analysis/InnerProductSpace/TwoDim.lean @@ -354,14 +354,16 @@ theorem inner_mul_inner_add_areaForm_mul_areaForm' (a x : E) : apply (o.basisRightAngleRotation a ha).ext intro i fin_cases i - · simp only [Fin.mk_zero, coe_basisRightAngleRotation, Matrix.cons_val_zero, LinearMap.add_apply, - LinearMap.smul_apply, innerₛₗ_apply, real_inner_self_eq_norm_sq, smul_eq_mul, - areaForm_apply_self, mul_zero, add_zero, Real.rpow_two, real_inner_comm] + · simp only [Fin.zero_eta, Fin.isValue, id_eq, coe_basisRightAngleRotation, Nat.succ_eq_add_one, + Nat.reduceAdd, Matrix.cons_val_zero, LinearMap.add_apply, LinearMap.smul_apply, innerₛₗ_apply, + real_inner_self_eq_norm_sq, smul_eq_mul, areaForm_apply_self, mul_zero, add_zero, + real_inner_comm] ring - · simp only [Fin.mk_one, coe_basisRightAngleRotation, Matrix.cons_val_one, Matrix.head_cons, - LinearMap.add_apply, LinearMap.smul_apply, innerₛₗ_apply, inner_rightAngleRotation_right, - areaForm_apply_self, neg_zero, smul_eq_mul, mul_zero, areaForm_rightAngleRotation_right, - real_inner_self_eq_norm_sq, zero_add, Real.rpow_two, mul_neg] + · simp only [Fin.mk_one, Fin.isValue, id_eq, coe_basisRightAngleRotation, Nat.succ_eq_add_one, + Nat.reduceAdd, Matrix.cons_val_one, Matrix.head_cons, LinearMap.add_apply, + LinearMap.smul_apply, innerₛₗ_apply, inner_rightAngleRotation_right, areaForm_apply_self, + neg_zero, smul_eq_mul, mul_zero, areaForm_rightAngleRotation_right, + real_inner_self_eq_norm_sq, zero_add, mul_neg] rw [o.areaForm_swap] ring @@ -381,15 +383,16 @@ theorem inner_mul_areaForm_sub' (a x : E) : ⟪a, x⟫ • ω a - ω a x • inn apply (o.basisRightAngleRotation a ha).ext intro i fin_cases i - · simp only [o.areaForm_swap a x, neg_smul, sub_neg_eq_add, Fin.mk_zero, - coe_basisRightAngleRotation, Matrix.cons_val_zero, LinearMap.add_apply, LinearMap.smul_apply, - areaForm_apply_self, smul_eq_mul, mul_zero, innerₛₗ_apply, real_inner_self_eq_norm_sq, - zero_add, Real.rpow_two] + · simp only [o.areaForm_swap a x, neg_smul, sub_neg_eq_add, Fin.zero_eta, Fin.isValue, id_eq, + coe_basisRightAngleRotation, Nat.succ_eq_add_one, Nat.reduceAdd, Matrix.cons_val_zero, + LinearMap.add_apply, LinearMap.smul_apply, areaForm_apply_self, smul_eq_mul, mul_zero, + innerₛₗ_apply, real_inner_self_eq_norm_sq, zero_add] ring - · simp only [Fin.mk_one, coe_basisRightAngleRotation, Matrix.cons_val_one, Matrix.head_cons, - LinearMap.sub_apply, LinearMap.smul_apply, areaForm_rightAngleRotation_right, - real_inner_self_eq_norm_sq, smul_eq_mul, innerₛₗ_apply, inner_rightAngleRotation_right, - areaForm_apply_self, neg_zero, mul_zero, sub_zero, Real.rpow_two, real_inner_comm] + · simp only [Fin.mk_one, Fin.isValue, id_eq, coe_basisRightAngleRotation, Nat.succ_eq_add_one, + Nat.reduceAdd, Matrix.cons_val_one, Matrix.head_cons, LinearMap.sub_apply, + LinearMap.smul_apply, areaForm_rightAngleRotation_right, real_inner_self_eq_norm_sq, + smul_eq_mul, innerₛₗ_apply, inner_rightAngleRotation_right, areaForm_apply_self, neg_zero, + mul_zero, sub_zero, real_inner_comm] ring /-- For vectors `a x y : E`, the identity `⟪a, x⟫ * ω a y - ω a x * ⟪a, y⟫ = ‖a‖ ^ 2 * ω x y`. -/ diff --git a/Mathlib/Data/List/Defs.lean b/Mathlib/Data/List/Defs.lean index 03674193faaec..4b86b7218cdfb 100644 --- a/Mathlib/Data/List/Defs.lean +++ b/Mathlib/Data/List/Defs.lean @@ -488,9 +488,17 @@ theorem length_mapAccumr₂ : end MapAccumr +/- #adaptation_note: this attribute should be removed after Mathlib moves to v4.15.0-rc1. -/ +set_option allowUnsafeReducibility true in +attribute [semireducible] Fin.foldr.loop + /-- All elements of `Fin n`, from `0` to `n-1`. The corresponding finset is `Finset.univ`. -/ -def finRange (n : ℕ) : List (Fin n) := - (range n).pmap Fin.mk fun _ => List.mem_range.1 +-- Note that we use `ofFn (fun x => x)` instead of `ofFn id` to avoid leaving `id` in the terms +-- for e.g. `Fintype (Fin n)`. +def finRange (n : Nat) : List (Fin n) := ofFn (fun x => x) + +-- Verify that `finRange` is semireducible. +example : finRange 3 = [0, 1, 2] := rfl section Deprecated diff --git a/Mathlib/Data/List/FinRange.lean b/Mathlib/Data/List/FinRange.lean index 62d8122999597..bf592cfe78937 100644 --- a/Mathlib/Data/List/FinRange.lean +++ b/Mathlib/Data/List/FinRange.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Kenny Lau, Kim Morrison, Alex Keizer -/ import Mathlib.Data.List.OfFn -import Mathlib.Data.List.Range import Batteries.Data.List.Perm +import Mathlib.Data.List.Nodup /-! # Lists of elements of `Fin n` @@ -21,10 +21,66 @@ namespace List variable {α : Type u} + +theorem finRange_eq_pmap_range (n : ℕ) : finRange n = (range n).pmap Fin.mk (by simp) := by + apply List.ext_getElem <;> simp [finRange] + +@[simp] +theorem finRange_zero : finRange 0 = [] := rfl + +@[simp] +theorem mem_finRange {n : ℕ} (a : Fin n) : a ∈ finRange n := by + rw [finRange_eq_pmap_range] + exact mem_pmap.2 + ⟨a.1, mem_range.2 a.2, by + cases a + rfl⟩ + +theorem nodup_finRange (n : ℕ) : (finRange n).Nodup := by + rw [finRange_eq_pmap_range] + exact (Pairwise.pmap (nodup_range n) _) fun _ _ _ _ => @Fin.ne_of_val_ne _ ⟨_, _⟩ ⟨_, _⟩ + +@[simp] +theorem length_finRange (n : ℕ) : (finRange n).length = n := by + simp [finRange] + +@[simp] +theorem finRange_eq_nil {n : ℕ} : finRange n = [] ↔ n = 0 := by + rw [← length_eq_zero, length_finRange] + +theorem pairwise_lt_finRange (n : ℕ) : Pairwise (· < ·) (finRange n) := by + rw [finRange_eq_pmap_range] + exact (List.pairwise_lt_range n).pmap (by simp) (by simp) + +theorem pairwise_le_finRange (n : ℕ) : Pairwise (· ≤ ·) (finRange n) := by + rw [finRange_eq_pmap_range] + exact (List.pairwise_le_range n).pmap (by simp) (by simp) + +@[simp] +theorem getElem_finRange {n : ℕ} {i : ℕ} (h) : + (finRange n)[i] = ⟨i, length_finRange n ▸ h⟩ := by + simp [finRange, getElem_range, getElem_pmap] + +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/10756): new theorem +theorem get_finRange {n : ℕ} {i : ℕ} (h) : + (finRange n).get ⟨i, h⟩ = ⟨i, length_finRange n ▸ h⟩ := by + simp + +@[deprecated (since := "2024-08-19")] alias nthLe_finRange := get_finRange + +@[simp] +theorem finRange_map_get (l : List α) : (finRange l.length).map l.get = l := + List.ext_get (by simp) (by simp) + +@[simp] theorem indexOf_finRange {k : ℕ} (i : Fin k) : (finRange k).indexOf i = i := by + have : (finRange k).indexOf i < (finRange k).length := indexOf_lt_length.mpr (by simp) + have h₁ : (finRange k).get ⟨(finRange k).indexOf i, this⟩ = i := indexOf_get this + have h₂ : (finRange k).get ⟨i, by simp⟩ = i := get_finRange _ + simpa using (Nodup.get_inj_iff (nodup_finRange k)).mp (Eq.trans h₁ h₂.symm) + @[simp] theorem map_coe_finRange (n : ℕ) : ((finRange n) : List (Fin n)).map (Fin.val) = List.range n := by - simp_rw [finRange, map_pmap, pmap_eq_map] - exact List.map_id _ + apply List.ext_getElem <;> simp theorem finRange_succ_eq_map (n : ℕ) : finRange n.succ = 0 :: (finRange n).map Fin.succ := by apply map_injective_iff.mpr Fin.val_injective @@ -45,7 +101,7 @@ theorem ofFn_eq_pmap {n} {f : Fin n → α} : exact ext_getElem (by simp) fun i hi1 hi2 => by simp [List.getElem_ofFn f i hi1] theorem ofFn_id (n) : ofFn id = finRange n := - ofFn_eq_pmap + rfl theorem ofFn_eq_map {n} {f : Fin n → α} : ofFn f = (finRange n).map f := by rw [← ofFn_id, map_ofFn, Function.comp_id] diff --git a/Mathlib/Data/List/Range.lean b/Mathlib/Data/List/Range.lean index 7cefc522da8e2..e133f0b507d6d 100644 --- a/Mathlib/Data/List/Range.lean +++ b/Mathlib/Data/List/Range.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Kenny Lau, Kim Morrison -/ import Mathlib.Data.List.Chain -import Mathlib.Data.List.Nodup /-! # Ranges of naturals as lists @@ -42,58 +41,10 @@ theorem chain_range_succ (r : ℕ → ℕ → Prop) (n a : ℕ) : rw [range_succ_eq_map, chain_cons, and_congr_right_iff, ← chain'_range_succ, range_succ_eq_map] exact fun _ => Iff.rfl -@[simp] -theorem finRange_zero : finRange 0 = [] := - rfl - -@[simp] -theorem mem_finRange {n : ℕ} (a : Fin n) : a ∈ finRange n := - mem_pmap.2 - ⟨a.1, mem_range.2 a.2, by - cases a - rfl⟩ - -theorem nodup_finRange (n : ℕ) : (finRange n).Nodup := - (Pairwise.pmap (nodup_range n) _) fun _ _ _ _ => @Fin.ne_of_val_ne _ ⟨_, _⟩ ⟨_, _⟩ - -@[simp] -theorem length_finRange (n : ℕ) : (finRange n).length = n := by - rw [finRange, length_pmap, length_range] - -@[simp] -theorem finRange_eq_nil {n : ℕ} : finRange n = [] ↔ n = 0 := by - rw [← length_eq_zero, length_finRange] - -theorem pairwise_lt_finRange (n : ℕ) : Pairwise (· < ·) (finRange n) := - (List.pairwise_lt_range n).pmap (by simp) (by simp) - -theorem pairwise_le_finRange (n : ℕ) : Pairwise (· ≤ ·) (finRange n) := - (List.pairwise_le_range n).pmap (by simp) (by simp) - -@[simp] -theorem getElem_finRange {n : ℕ} {i : ℕ} (h) : - (finRange n)[i] = ⟨i, length_finRange n ▸ h⟩ := by - simp [finRange, getElem_range, getElem_pmap] - --- Porting note (https://github.com/leanprover-community/mathlib4/issues/10756): new theorem -theorem get_finRange {n : ℕ} {i : ℕ} (h) : - (finRange n).get ⟨i, h⟩ = ⟨i, length_finRange n ▸ h⟩ := by - simp - @[deprecated (since := "2024-08-19")] alias nthLe_range' := get_range' @[deprecated (since := "2024-08-19")] alias nthLe_range'_1 := getElem_range'_1 @[deprecated (since := "2024-08-19")] alias nthLe_range := get_range -@[deprecated (since := "2024-08-19")] alias nthLe_finRange := get_finRange - -@[simp] -theorem finRange_map_get (l : List α) : (finRange l.length).map l.get = l := - List.ext_get (by simp) (by simp) -@[simp] theorem indexOf_finRange {k : ℕ} (i : Fin k) : (finRange k).indexOf i = i := by - have : (finRange k).indexOf i < (finRange k).length := indexOf_lt_length.mpr (by simp) - have h₁ : (finRange k).get ⟨(finRange k).indexOf i, this⟩ = i := indexOf_get this - have h₂ : (finRange k).get ⟨i, by simp⟩ = i := get_finRange _ - simpa using (Nodup.get_inj_iff (nodup_finRange k)).mp (Eq.trans h₁ h₂.symm) section Ranges diff --git a/Mathlib/Data/List/Sublists.lean b/Mathlib/Data/List/Sublists.lean index cc89adca8c9e8..9a2c8391406a8 100644 --- a/Mathlib/Data/List/Sublists.lean +++ b/Mathlib/Data/List/Sublists.lean @@ -4,8 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Mathlib.Data.Nat.Choose.Basic -import Mathlib.Data.List.Range +import Mathlib.Data.List.FinRange import Mathlib.Data.List.Perm.Basic +import Mathlib.Data.List.Lex /-! # sublists diff --git a/Mathlib/LinearAlgebra/AffineSpace/Independent.lean b/Mathlib/LinearAlgebra/AffineSpace/Independent.lean index 9476da73472ed..308268392107c 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Independent.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Independent.lean @@ -618,7 +618,7 @@ theorem affineIndependent_of_ne {p₁ p₂ : P} (h : p₁ ≠ p₂) : AffineInde ext fin_cases i · simp at hi - · simp only [Fin.val_one] + · simp haveI : Unique { x // x ≠ (0 : Fin 2) } := ⟨⟨i₁⟩, he'⟩ apply linearIndependent_unique rw [he' default] diff --git a/Mathlib/NumberTheory/Modular.lean b/Mathlib/NumberTheory/Modular.lean index 7eb5dbb5ffec4..59b083317b35b 100644 --- a/Mathlib/NumberTheory/Modular.lean +++ b/Mathlib/NumberTheory/Modular.lean @@ -207,17 +207,19 @@ theorem tendsto_lcRow0 {cd : Fin 2 → ℤ} (hcd : IsCoprime (cd 0) (cd 1)) : ext ⟨g, rfl⟩ i j : 3 fin_cases i <;> [fin_cases j; skip] -- the following are proved by `simp`, but it is replaced by `simp only` to avoid timeouts. - · simp only [mB, mulVec, dotProduct, Fin.sum_univ_two, coe_matrix_coe, - Int.coe_castRingHom, lcRow0_apply, Function.comp_apply, cons_val_zero, lcRow0Extend_apply, + · simp only [Fin.isValue, Int.cast_one, map_apply_coe, RingHom.mapMatrix_apply, + Int.coe_castRingHom, lcRow0_apply, map_apply, Fin.zero_eta, id_eq, Function.comp_apply, + of_apply, cons_val', cons_val_zero, empty_val', cons_val_fin_one, lcRow0Extend_apply, LinearMap.GeneralLinearGroup.coeFn_generalLinearEquiv, GeneralLinearGroup.coe_toLinear, - val_planeConformalMatrix, neg_neg, mulVecLin_apply, cons_val_one, head_cons, of_apply, - Fin.mk_zero, Fin.mk_one] + val_planeConformalMatrix, neg_neg, mulVecLin_apply, mulVec, dotProduct, Fin.sum_univ_two, + cons_val_one, head_cons, mB, f₁] · convert congr_arg (fun n : ℤ => (-n : ℝ)) g.det_coe.symm using 1 - simp only [f₁, mulVec, dotProduct, Fin.sum_univ_two, Matrix.det_fin_two, Function.comp_apply, - Subtype.coe_mk, lcRow0Extend_apply, cons_val_zero, + simp only [Fin.zero_eta, id_eq, Function.comp_apply, lcRow0Extend_apply, cons_val_zero, LinearMap.GeneralLinearGroup.coeFn_generalLinearEquiv, GeneralLinearGroup.coe_toLinear, - val_planeConformalMatrix, mulVecLin_apply, cons_val_one, head_cons, map_apply, neg_mul, - Int.cast_sub, Int.cast_mul, neg_sub, of_apply, Fin.mk_zero, Fin.mk_one] + mulVecLin_apply, mulVec, dotProduct, det_fin_two, f₁] + simp only [Fin.isValue, Fin.mk_one, val_planeConformalMatrix, neg_neg, of_apply, cons_val', + empty_val', cons_val_fin_one, cons_val_one, head_fin_const, map_apply, Fin.sum_univ_two, + cons_val_zero, neg_mul, head_cons, Int.cast_sub, Int.cast_mul, neg_sub] ring · rfl diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean index 84f4e2aa14e91..3b10a1368dea7 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean @@ -585,17 +585,17 @@ theorem stdBasis_repr_eq_matrixToStdBasis_mul (x : (K →+* ℂ) → ℂ) | inr c => rcases c with ⟨w, j⟩ fin_cases j - · simp_rw [Fin.mk_zero, stdBasis_apply_ofIsComplex_fst, fromBlocks_apply₂₁, - fromBlocks_apply₂₂, Matrix.zero_apply, submatrix_apply, - blockDiagonal_apply, Prod.swap_prod_mk, ite_mul, zero_mul, sum_const_zero, zero_add, - sum_add_distrib, sum_ite_eq, mem_univ, ite_true, of_apply, cons_val', cons_val_zero, - cons_val_one, head_cons, ← hx (embedding w), re_eq_add_conj] + · simp only [Fin.zero_eta, Fin.isValue, id_eq, stdBasis_apply_ofIsComplex_fst, re_eq_add_conj, + mul_neg, fromBlocks_apply₂₁, zero_apply, zero_mul, sum_const_zero, fromBlocks_apply₂₂, + submatrix_apply, Prod.swap_prod_mk, blockDiagonal_apply, of_apply, cons_val', cons_val_zero, + empty_val', cons_val_fin_one, ite_mul, cons_val_one, head_cons, sum_add_distrib, sum_ite_eq, + mem_univ, ↓reduceIte, ← hx (embedding w), zero_add] field_simp - · simp_rw [Fin.mk_one, stdBasis_apply_ofIsComplex_snd, fromBlocks_apply₂₁, - fromBlocks_apply₂₂, Matrix.zero_apply, submatrix_apply, blockDiagonal_apply, - Prod.swap_prod_mk, ite_mul, zero_mul, sum_const_zero, zero_add, sum_add_distrib, sum_ite_eq, - mem_univ, ite_true, of_apply, cons_val', cons_val_zero, cons_val_one, head_cons, - ← hx (embedding w), im_eq_sub_conj] + · simp only [Fin.mk_one, Fin.isValue, id_eq, stdBasis_apply_ofIsComplex_snd, im_eq_sub_conj, + mul_neg, fromBlocks_apply₂₁, zero_apply, zero_mul, sum_const_zero, fromBlocks_apply₂₂, + submatrix_apply, Prod.swap_prod_mk, blockDiagonal_apply, of_apply, cons_val', cons_val_zero, + empty_val', cons_val_fin_one, cons_val_one, head_fin_const, ite_mul, neg_mul, head_cons, + sum_add_distrib, sum_ite_eq, mem_univ, ↓reduceIte, ← hx (embedding w), zero_add] ring_nf; field_simp end stdBasis diff --git a/Mathlib/Order/RelSeries.lean b/Mathlib/Order/RelSeries.lean index b2c41395b294b..e385b759958c2 100644 --- a/Mathlib/Order/RelSeries.lean +++ b/Mathlib/Order/RelSeries.lean @@ -112,9 +112,9 @@ corresponds to each other. -/ protected def Equiv : RelSeries r ≃ {x : List α | x ≠ [] ∧ x.Chain' r} where toFun x := ⟨_, x.toList_ne_nil, x.toList_chain'⟩ invFun x := fromListChain' _ x.2.1 x.2.2 - left_inv x := ext (by simp [toList]) <| by ext; apply List.get_ofFn + left_inv x := ext (by simp [toList]) <| by ext; dsimp; apply List.get_ofFn right_inv x := by - refine Subtype.ext (List.ext_get ?_ fun n hn1 _ => List.get_ofFn _ _) + refine Subtype.ext (List.ext_get ?_ fun n hn1 _ => by dsimp; apply List.get_ofFn) have := Nat.succ_pred_eq_of_pos <| List.length_pos.mpr x.2.1 simp_all [toList] diff --git a/Mathlib/RingTheory/Complex.lean b/Mathlib/RingTheory/Complex.lean index 68be8e6b337a3..1f3b98e14ae26 100644 --- a/Mathlib/RingTheory/Complex.lean +++ b/Mathlib/RingTheory/Complex.lean @@ -18,11 +18,11 @@ theorem Algebra.leftMulMatrix_complex (z : ℂ) : rw [Algebra.leftMulMatrix_eq_repr_mul, Complex.coe_basisOneI_repr, Complex.coe_basisOneI, mul_re, mul_im, Matrix.of_apply] fin_cases j - · simp only [Fin.mk_zero, Matrix.cons_val_zero, one_re, mul_one, one_im, mul_zero, sub_zero, - zero_add] + · simp only [Fin.zero_eta, id_eq, Matrix.cons_val_zero, one_re, mul_one, one_im, mul_zero, + sub_zero, zero_add, Matrix.cons_val_fin_one] fin_cases i <;> rfl - · simp only [Fin.mk_one, Matrix.cons_val_one, Matrix.head_cons, I_re, mul_zero, I_im, mul_one, - zero_sub, add_zero] + · simp only [Fin.mk_one, id_eq, Matrix.cons_val_one, Matrix.head_cons, I_re, mul_zero, I_im, + mul_one, zero_sub, add_zero, Matrix.cons_val_fin_one] fin_cases i <;> rfl theorem Algebra.trace_complex_apply (z : ℂ) : Algebra.trace ℝ ℂ z = 2 * z.re := by diff --git a/Mathlib/SetTheory/Cardinal/Subfield.lean b/Mathlib/SetTheory/Cardinal/Subfield.lean index 604447735f1cc..37f45168df1dd 100644 --- a/Mathlib/SetTheory/Cardinal/Subfield.lean +++ b/Mathlib/SetTheory/Cardinal/Subfield.lean @@ -76,7 +76,7 @@ lemma cardinalMk_closure_le_max : #(closure s) ≤ max #s ℵ₀ := exact (add_lt_aleph0 this h).le · rw [max_eq_left h, add_eq_right h (this.le.trans h), max_eq_left h] rintro (n|_) - · fin_cases n <;> infer_instance + · fin_cases n <;> (dsimp only [id_eq]; infer_instance) infer_instance @[deprecated (since := "2024-11-10")] alias cardinal_mk_closure_le_max := cardinalMk_closure_le_max From 3e56ff8fc708e700c810326438a65f57315ec0e9 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Mon, 25 Nov 2024 15:15:47 +0000 Subject: [PATCH 290/829] feat(CategoryTheory): limits of eventually constant functors from cofiltered categories (#19021) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit An eventually constant functor from a cofiltered category admits a limit. Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Filtered/Basic.lean | 20 ++ .../Constructions/EventuallyConstant.lean | 260 ++++++++++++++++++ 3 files changed, 281 insertions(+) create mode 100644 Mathlib/CategoryTheory/Limits/Constructions/EventuallyConstant.lean diff --git a/Mathlib.lean b/Mathlib.lean index 39dccdce392ff..02bb3ae68c288 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1770,6 +1770,7 @@ import Mathlib.CategoryTheory.Limits.Connected import Mathlib.CategoryTheory.Limits.Constructions.BinaryProducts import Mathlib.CategoryTheory.Limits.Constructions.EpiMono import Mathlib.CategoryTheory.Limits.Constructions.Equalizers +import Mathlib.CategoryTheory.Limits.Constructions.EventuallyConstant import Mathlib.CategoryTheory.Limits.Constructions.Filtered import Mathlib.CategoryTheory.Limits.Constructions.FiniteProductsOfBinaryProducts import Mathlib.CategoryTheory.Limits.Constructions.LimitsOfProductsAndEqualizers diff --git a/Mathlib/CategoryTheory/Filtered/Basic.lean b/Mathlib/CategoryTheory/Filtered/Basic.lean index 28de77ac9e3bb..584411d3326a4 100644 --- a/Mathlib/CategoryTheory/Filtered/Basic.lean +++ b/Mathlib/CategoryTheory/Filtered/Basic.lean @@ -616,6 +616,26 @@ theorem _root_.CategoryTheory.Functor.ranges_directed (F : C ⥤ Type*) (j : C) let ⟨l, li, lk, e⟩ := cospan ij kj refine ⟨⟨l, lk ≫ kj⟩, e ▸ ?_, ?_⟩ <;> simp_rw [F.map_comp] <;> apply Set.range_comp_subset_range +/-- Given a "bowtie" of morphisms +``` + k₁ k₂ + |\ /| + | \/ | + | /\ | + |/ \∣ + vv vv + j₁ j₂ +``` +in a cofiltered category, we can construct an object `s` and two morphisms +from `s` to `k₁` and `k₂`, making the resulting squares commute. +-/ +theorem bowtie {j₁ j₂ k₁ k₂ : C} (f₁ : k₁ ⟶ j₁) (g₁ : k₂ ⟶ j₁) (f₂ : k₁ ⟶ j₂) (g₂ : k₂ ⟶ j₂) : + ∃ (s : C) (α : s ⟶ k₁) (β : s ⟶ k₂), α ≫ f₁ = β ≫ g₁ ∧ α ≫ f₂ = β ≫ g₂ := by + obtain ⟨t, k₁t, k₂t, ht⟩ := cospan f₁ g₁ + obtain ⟨s, ts, hs⟩ := IsCofilteredOrEmpty.cone_maps (k₁t ≫ f₂) (k₂t ≫ g₂) + exact ⟨s, ts ≫ k₁t, ts ≫ k₂t, by simp only [Category.assoc, ht], + by simp only [Category.assoc, hs]⟩ + end AllowEmpty end IsCofiltered diff --git a/Mathlib/CategoryTheory/Limits/Constructions/EventuallyConstant.lean b/Mathlib/CategoryTheory/Limits/Constructions/EventuallyConstant.lean new file mode 100644 index 0000000000000..9ae75d85d62ae --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Constructions/EventuallyConstant.lean @@ -0,0 +1,260 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.CategoryTheory.Filtered.Basic +import Mathlib.CategoryTheory.Limits.HasLimits + +/-! +# Limits of eventually constant functors + +If `F : J ⥤ C` is a functor from a cofiltered category, and `j : J`, +we introduce a property `F.IsEventuallyConstantTo j` which says +that for any `f : i ⟶ j`, the induced morphism `F.map f` is an isomorphism. +Under this assumption, it is shown that `F` admits `F.obj j` as a limit +(`Functor.IsEventuallyConstantTo.isLimitCone`). + +A typeclass `Cofiltered.IsEventuallyConstant` is also introduced, and +the dual results for filtered categories and colimits are also obtained. + +-/ + +namespace CategoryTheory + +open Category Limits + +variable {J C : Type*} [Category J] [Category C] (F : J ⥤ C) + +namespace Functor + +/-- A functor `F : J ⥤ C` is eventually constant to `j : J` if +for any map `f : i ⟶ j`, the induced morphism `F.map f` is an isomorphism. +If `J` is cofiltered, this implies `F` has a limit. -/ +def IsEventuallyConstantTo (j : J) : Prop := + ∀ ⦃i : J⦄ (f : i ⟶ j), IsIso (F.map f) + +/-- A functor `F : J ⥤ C` is eventually constant from `i : J` if +for any map `f : i ⟶ j`, the induced morphism `F.map f` is an isomorphism. +If `J` is filtered, this implies `F` has a colimit. -/ +def IsEventuallyConstantFrom (i : J) : Prop := + ∀ ⦃j : J⦄ (f : i ⟶ j), IsIso (F.map f) + +namespace IsEventuallyConstantTo + +variable {F} {i₀ : J} (h : F.IsEventuallyConstantTo i₀) + +include h + +lemma isIso_map {i j : J} (φ : i ⟶ j) (π : j ⟶ i₀) : IsIso (F.map φ) := by + have := h π + have := h (φ ≫ π) + exact IsIso.of_isIso_fac_right (F.map_comp φ π).symm + +lemma precomp {j : J} (f : j ⟶ i₀) : F.IsEventuallyConstantTo j := + fun _ φ ↦ h.isIso_map φ f + +section + +variable {i j : J} (φ : i ⟶ j) (hφ : Nonempty (j ⟶ i₀)) + +/-- The isomorphism `F.obj i ≅ F.obj j` induced by `φ : i ⟶ j`, +when `h : F.IsEventuallyConstantTo i₀` and there exists a map `j ⟶ i₀`. -/ +@[simps! hom] +noncomputable def isoMap : F.obj i ≅ F.obj j := + have := h.isIso_map φ hφ.some + asIso (F.map φ) + +@[reassoc (attr := simp)] +lemma isoMap_hom_inv_id : F.map φ ≫ (h.isoMap φ hφ).inv = 𝟙 _ := + (h.isoMap φ hφ).hom_inv_id + +@[reassoc (attr := simp)] +lemma isoMap_inv_hom_id : (h.isoMap φ hφ).inv ≫ F.map φ = 𝟙 _ := + (h.isoMap φ hφ).inv_hom_id + +end + +variable [IsCofiltered J] +open IsCofiltered + +/-- Auxiliary definition for `IsEventuallyConstantTo.cone`. -/ +noncomputable def coneπApp (j : J) : F.obj i₀ ⟶ F.obj j := + (h.isoMap (minToLeft i₀ j) ⟨𝟙 _⟩).inv ≫ F.map (minToRight i₀ j) + +lemma coneπApp_eq (j j' : J) (α : j' ⟶ i₀) (β : j' ⟶ j) : + h.coneπApp j = (h.isoMap α ⟨𝟙 _⟩).inv ≫ F.map β := by + obtain ⟨s, γ, δ, h₁, h₂⟩ := IsCofiltered.bowtie + (IsCofiltered.minToRight i₀ j) β (IsCofiltered.minToLeft i₀ j) α + dsimp [coneπApp] + rw [← cancel_epi ((h.isoMap α ⟨𝟙 _⟩).hom), isoMap_hom, isoMap_hom_inv_id_assoc, + ← cancel_epi (h.isoMap δ ⟨α⟩).hom, isoMap_hom, + ← F.map_comp δ β, ← h₁, F.map_comp, ← F.map_comp_assoc, ← h₂, F.map_comp_assoc, + isoMap_hom_inv_id_assoc] + +@[simp] +lemma coneπApp_eq_id : h.coneπApp i₀ = 𝟙 _ := by + rw [h.coneπApp_eq i₀ i₀ (𝟙 _) (𝟙 _), h.isoMap_inv_hom_id] + +/-- Given `h : F.IsEventuallyConstantTo i₀`, this is the (limit) cone for `F` whose +point is `F.obj i₀`. -/ +@[simps] +noncomputable def cone : Cone F where + pt := F.obj i₀ + π := + { app := h.coneπApp + naturality := fun j j' φ ↦ by + dsimp + rw [id_comp] + let i := IsCofiltered.min i₀ j + let α : i ⟶ i₀ := IsCofiltered.minToLeft _ _ + let β : i ⟶ j := IsCofiltered.minToRight _ _ + rw [h.coneπApp_eq j _ α β, assoc, h.coneπApp_eq j' _ α (β ≫ φ), map_comp] } + +/-- When `h : F.IsEventuallyConstantTo i₀`, the limit of `F` exists and is `F.obj i₀`. -/ +def isLimitCone : IsLimit h.cone where + lift s := s.π.app i₀ + fac s j := by + dsimp [coneπApp] + rw [← s.w (IsCofiltered.minToLeft i₀ j), ← s.w (IsCofiltered.minToRight i₀ j), assoc, + isoMap_hom_inv_id_assoc] + uniq s m hm := by simp only [← hm i₀, cone_π_app, coneπApp_eq_id, cone_pt, comp_id] + +lemma hasLimit : HasLimit F := ⟨_, h.isLimitCone⟩ + +lemma isIso_π_of_isLimit {c : Cone F} (hc : IsLimit c) : + IsIso (c.π.app i₀) := by + simp only [← IsLimit.conePointUniqueUpToIso_hom_comp hc h.isLimitCone i₀, + cone_π_app, coneπApp_eq_id, cone_pt, comp_id] + infer_instance + +/-- More general version of `isIso_π_of_isLimit`. -/ +lemma isIso_π_of_isLimit' {c : Cone F} (hc : IsLimit c) (j : J) (π : j ⟶ i₀) : + IsIso (c.π.app j) := + (h.precomp π).isIso_π_of_isLimit hc + +end IsEventuallyConstantTo + +namespace IsEventuallyConstantFrom + +variable {F} {i₀ : J} (h : F.IsEventuallyConstantFrom i₀) + +include h + +lemma isIso_map {i j : J} (φ : i ⟶ j) (ι : i₀ ⟶ i) : IsIso (F.map φ) := by + have := h ι + have := h (ι ≫ φ) + exact IsIso.of_isIso_fac_left (F.map_comp ι φ).symm + +lemma postcomp {j : J} (f : i₀ ⟶ j) : F.IsEventuallyConstantFrom j := + fun _ φ ↦ h.isIso_map φ f + +section + +variable {i j : J} (φ : i ⟶ j) (hφ : Nonempty (i₀ ⟶ i)) + +/-- The isomorphism `F.obj i ≅ F.obj j` induced by `φ : i ⟶ j`, +when `h : F.IsEventuallyConstantFrom i₀` and there exists a map `i₀ ⟶ i`. -/ +@[simps! hom] +noncomputable def isoMap : F.obj i ≅ F.obj j := + have := h.isIso_map φ hφ.some + asIso (F.map φ) + +@[reassoc (attr := simp)] +lemma isoMap_hom_inv_id : F.map φ ≫ (h.isoMap φ hφ).inv = 𝟙 _ := + (h.isoMap φ hφ).hom_inv_id + +@[reassoc (attr := simp)] +lemma isoMap_inv_hom_id : (h.isoMap φ hφ).inv ≫ F.map φ = 𝟙 _ := + (h.isoMap φ hφ).inv_hom_id + +end + +variable [IsFiltered J] +open IsFiltered + +/-- Auxiliary definition for `IsEventuallyConstantFrom.cocone`. -/ +noncomputable def coconeιApp (j : J) : F.obj j ⟶ F.obj i₀ := + F.map (rightToMax i₀ j) ≫ (h.isoMap (leftToMax i₀ j) ⟨𝟙 _⟩).inv + +lemma coconeιApp_eq (j j' : J) (α : j ⟶ j') (β : i₀ ⟶ j') : + h.coconeιApp j = F.map α ≫ (h.isoMap β ⟨𝟙 _⟩).inv := by + obtain ⟨s, γ, δ, h₁, h₂⟩ := IsFiltered.bowtie + (IsFiltered.leftToMax i₀ j) β (IsFiltered.rightToMax i₀ j) α + dsimp [coconeιApp] + rw [← cancel_mono ((h.isoMap β ⟨𝟙 _⟩).hom), assoc, assoc, isoMap_hom, isoMap_inv_hom_id, + comp_id, ← cancel_mono (h.isoMap δ ⟨β⟩).hom, isoMap_hom, assoc, assoc, ← F.map_comp α δ, + ← h₂, F.map_comp, ← F.map_comp β δ, ← h₁, F.map_comp, isoMap_inv_hom_id_assoc] + +@[simp] +lemma coconeιApp_eq_id : h.coconeιApp i₀ = 𝟙 _ := by + rw [h.coconeιApp_eq i₀ i₀ (𝟙 _) (𝟙 _), h.isoMap_hom_inv_id] + +/-- Given `h : F.IsEventuallyConstantFrom i₀`, this is the (limit) cocone for `F` whose +point is `F.obj i₀`. -/ +@[simps] +noncomputable def cocone : Cocone F where + pt := F.obj i₀ + ι := + { app := h.coconeιApp + naturality := fun j j' φ ↦ by + dsimp + rw [comp_id] + let i := IsFiltered.max i₀ j' + let α : i₀ ⟶ i := IsFiltered.leftToMax _ _ + let β : j' ⟶ i := IsFiltered.rightToMax _ _ + rw [h.coconeιApp_eq j' _ β α, h.coconeιApp_eq j _ (φ ≫ β) α, map_comp, assoc] } + +/-- When `h : F.IsEventuallyConstantFrom i₀`, the colimit of `F` exists and is `F.obj i₀`. -/ +def isColimitCocone : IsColimit h.cocone where + desc s := s.ι.app i₀ + fac s j := by + dsimp [coconeιApp] + rw [← s.w (IsFiltered.rightToMax i₀ j), ← s.w (IsFiltered.leftToMax i₀ j), assoc, + isoMap_inv_hom_id_assoc] + uniq s m hm := by simp only [← hm i₀, cocone_ι_app, coconeιApp_eq_id, id_comp] + +lemma hasColimit : HasColimit F := ⟨_, h.isColimitCocone⟩ + +lemma isIso_ι_of_isColimit {c : Cocone F} (hc : IsColimit c) : + IsIso (c.ι.app i₀) := by + simp only [← IsColimit.comp_coconePointUniqueUpToIso_inv hc h.isColimitCocone i₀, + cocone_ι_app, coconeιApp_eq_id, id_comp] + infer_instance + +/-- More general version of `isIso_ι_of_isColimit`. -/ +lemma isIso_ι_of_isColimit' {c : Cocone F} (hc : IsColimit c) (j : J) (ι : i₀ ⟶ j) : + IsIso (c.ι.app j) := + (h.postcomp ι).isIso_ι_of_isColimit hc + +end IsEventuallyConstantFrom + +end Functor + +namespace IsCofiltered + +/-- A functor `F : J ⥤ C` from a cofiltered category is eventually constant if there +exists `j : J`, such that for any `f : i ⟶ j`, the induced map `F.map f` is an isomorphism. -/ +class IsEventuallyConstant : Prop where + exists_isEventuallyConstantTo : ∃ (j : J), F.IsEventuallyConstantTo j + +instance [hF : IsEventuallyConstant F] [IsCofiltered J] : HasLimit F := by + obtain ⟨j, h⟩ := hF.exists_isEventuallyConstantTo + exact h.hasLimit + +end IsCofiltered + +namespace IsFiltered + +/-- A functor `F : J ⥤ C` from a filtered category is eventually constant if there +exists `i : J`, such that for any `f : i ⟶ j`, the induced map `F.map f` is an isomorphism. -/ +class IsEventuallyConstant : Prop where + exists_isEventuallyConstantFrom : ∃ (i : J), F.IsEventuallyConstantFrom i + +instance [hF : IsEventuallyConstant F] [IsFiltered J] : HasColimit F := by + obtain ⟨j, h⟩ := hF.exists_isEventuallyConstantFrom + exact h.hasColimit + +end IsFiltered + +end CategoryTheory From b9e31e9b11f73ac54f481c8fe732f94d825895a4 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Mon, 25 Nov 2024 15:15:49 +0000 Subject: [PATCH 291/829] feat(CategoryTheory): final functors and preservation/reflection/creation of colimits (#19452) Co-authored-by: Markus Himmel --- Mathlib/CategoryTheory/Limits/Final.lean | 126 ++++++++++++++++++++++- 1 file changed, 125 insertions(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Limits/Final.lean b/Mathlib/CategoryTheory/Limits/Final.lean index 69444f0842058..6719a13bc510a 100644 --- a/Mathlib/CategoryTheory/Limits/Final.lean +++ b/Mathlib/CategoryTheory/Limits/Final.lean @@ -65,7 +65,7 @@ Dualise condition 3 above and the implications 2 ⇒ 3 and 3 ⇒ 1 to initial fu noncomputable section -universe v v₁ v₂ v₃ u₁ u₂ u₃ +universe v v₁ v₂ v₃ v₄ u₁ u₂ u₃ u₄ namespace CategoryTheory @@ -305,6 +305,27 @@ def colimitCoconeComp (t : ColimitCocone G) : ColimitCocone (F ⋙ G) where instance (priority := 100) comp_hasColimit [HasColimit G] : HasColimit (F ⋙ G) := HasColimit.mk (colimitCoconeComp F (getColimitCocone G)) +instance (priority := 100) comp_preservesColimit {B : Type u₄} [Category.{v₄} B] {H : E ⥤ B} + [PreservesColimit G H] : PreservesColimit (F ⋙ G) H where + preserves {c} hc := by + refine ⟨isColimitExtendCoconeEquiv (G := G ⋙ H) F (H.mapCocone c) ?_⟩ + let hc' := isColimitOfPreserves H ((isColimitExtendCoconeEquiv F c).symm hc) + exact IsColimit.ofIsoColimit hc' (Cocones.ext (Iso.refl _) (by simp)) + +instance (priority := 100) comp_reflectsColimit {B : Type u₄} [Category.{v₄} B] {H : E ⥤ B} + [ReflectsColimit G H] : ReflectsColimit (F ⋙ G) H where + reflects {c} hc := by + refine ⟨isColimitExtendCoconeEquiv F _ (isColimitOfReflects H ?_)⟩ + let hc' := (isColimitExtendCoconeEquiv (G := G ⋙ H) F _).symm hc + exact IsColimit.ofIsoColimit hc' (Cocones.ext (Iso.refl _) (by simp)) + +instance (priority := 100) compCreatesColimit {B : Type u₄} [Category.{v₄} B] {H : E ⥤ B} + [CreatesColimit G H] : CreatesColimit (F ⋙ G) H where + lifts {c} hc := by + refine ⟨(liftColimit ((isColimitExtendCoconeEquiv F (G := G ⋙ H) _).symm hc)).whisker F, ?_⟩ + let i := liftedColimitMapsToOriginal ((isColimitExtendCoconeEquiv F (G := G ⋙ H) _).symm hc) + exact (Cocones.whiskering F).mapIso i ≪≫ ((coconesEquiv F (G ⋙ H)).unitIso.app _).symm + instance colimit_pre_isIso [HasColimit G] : IsIso (colimit.pre G F) := by rw [colimit.pre_eq (colimitCoconeComp F (getColimitCocone G)) (getColimitCocone G)] erw [IsColimit.desc_self] @@ -361,10 +382,51 @@ We can't make this an instance, because `F` is not determined by the goal. theorem hasColimit_of_comp [HasColimit (F ⋙ G)] : HasColimit G := HasColimit.mk (colimitCoconeOfComp F (getColimitCocone (F ⋙ G))) +theorem preservesColimit_of_comp {B : Type u₄} [Category.{v₄} B] {H : E ⥤ B} + [PreservesColimit (F ⋙ G) H] : PreservesColimit G H where + preserves {c} hc := by + refine ⟨isColimitWhiskerEquiv F _ ?_⟩ + let hc' := isColimitOfPreserves H ((isColimitWhiskerEquiv F _).symm hc) + exact IsColimit.ofIsoColimit hc' (Cocones.ext (Iso.refl _) (by simp)) + +theorem reflectsColimit_of_comp {B : Type u₄} [Category.{v₄} B] {H : E ⥤ B} + [ReflectsColimit (F ⋙ G) H] : ReflectsColimit G H where + reflects {c} hc := by + refine ⟨isColimitWhiskerEquiv F _ (isColimitOfReflects H ?_)⟩ + let hc' := (isColimitWhiskerEquiv F _).symm hc + exact IsColimit.ofIsoColimit hc' (Cocones.ext (Iso.refl _) (by simp)) + +/-- If `F` is final and `F ⋙ G` creates colimits of `H`, then so does `G`. -/ +def createsColimitOfComp {B : Type u₄} [Category.{v₄} B] {H : E ⥤ B} + [CreatesColimit (F ⋙ G) H] : CreatesColimit G H where + reflects := (reflectsColimit_of_comp F).reflects + lifts {c} hc := by + refine ⟨(extendCocone (F := F)).obj (liftColimit ((isColimitWhiskerEquiv F _).symm hc)), ?_⟩ + let i := liftedColimitMapsToOriginal (K := (F ⋙ G)) ((isColimitWhiskerEquiv F _).symm hc) + refine ?_ ≪≫ ((extendCocone (F := F)).mapIso i) ≪≫ ((coconesEquiv F (G ⋙ H)).counitIso.app _) + exact Cocones.ext (Iso.refl _) + include F in theorem hasColimitsOfShape_of_final [HasColimitsOfShape C E] : HasColimitsOfShape D E where has_colimit := fun _ => hasColimit_of_comp F +include F in +theorem preservesColimitsOfShape_of_final {B : Type u₄} [Category.{v₄} B] (H : E ⥤ B) + [PreservesColimitsOfShape C H] : PreservesColimitsOfShape D H where + preservesColimit := preservesColimit_of_comp F + +include F in +theorem reflectsColimitsOfShape_of_final {B : Type u₄} [Category.{v₄} B] (H : E ⥤ B) + [ReflectsColimitsOfShape C H] : ReflectsColimitsOfShape D H where + reflectsColimit := reflectsColimit_of_comp F + +include F in +/-- If `H` creates colimits of shape `C` and `F : C ⥤ D` is final, then `H` creates colimits of +shape `D`. -/ +def createsColimitsOfShapeOfFinal {B : Type u₄} [Category.{v₄} B] (H : E ⥤ B) + [CreatesColimitsOfShape C H] : CreatesColimitsOfShape D H where + CreatesColimit := createsColimitOfComp F + end Final end ArbitraryUniverse @@ -592,6 +654,27 @@ def limitConeComp (t : LimitCone G) : LimitCone (F ⋙ G) where instance (priority := 100) comp_hasLimit [HasLimit G] : HasLimit (F ⋙ G) := HasLimit.mk (limitConeComp F (getLimitCone G)) +instance (priority := 100) comp_preservesLimit {B : Type u₄} [Category.{v₄} B] {H : E ⥤ B} + [PreservesLimit G H] : PreservesLimit (F ⋙ G) H where + preserves {c} hc := by + refine ⟨isLimitExtendConeEquiv (G := G ⋙ H) F (H.mapCone c) ?_⟩ + let hc' := isLimitOfPreserves H ((isLimitExtendConeEquiv F c).symm hc) + exact IsLimit.ofIsoLimit hc' (Cones.ext (Iso.refl _) (by simp)) + +instance (priority := 100) comp_reflectsLimit {B : Type u₄} [Category.{v₄} B] {H : E ⥤ B} + [ReflectsLimit G H] : ReflectsLimit (F ⋙ G) H where + reflects {c} hc := by + refine ⟨isLimitExtendConeEquiv F _ (isLimitOfReflects H ?_)⟩ + let hc' := (isLimitExtendConeEquiv (G := G ⋙ H) F _).symm hc + exact IsLimit.ofIsoLimit hc' (Cones.ext (Iso.refl _) (by simp)) + +instance (priority := 100) compCreatesLimit {B : Type u₄} [Category.{v₄} B] {H : E ⥤ B} + [CreatesLimit G H] : CreatesLimit (F ⋙ G) H where + lifts {c} hc := by + refine ⟨(liftLimit ((isLimitExtendConeEquiv F (G := G ⋙ H) _).symm hc)).whisker F, ?_⟩ + let i := liftedLimitMapsToOriginal ((isLimitExtendConeEquiv F (G := G ⋙ H) _).symm hc) + exact (Cones.whiskering F).mapIso i ≪≫ ((conesEquiv F (G ⋙ H)).unitIso.app _).symm + instance limit_pre_isIso [HasLimit G] : IsIso (limit.pre G F) := by rw [limit.pre_eq (limitConeComp F (getLimitCone G)) (getLimitCone G)] erw [IsLimit.lift_self] @@ -637,10 +720,51 @@ We can't make this an instance, because `F` is not determined by the goal. theorem hasLimit_of_comp [HasLimit (F ⋙ G)] : HasLimit G := HasLimit.mk (limitConeOfComp F (getLimitCone (F ⋙ G))) +theorem preservesLimit_of_comp {B : Type u₄} [Category.{v₄} B] {H : E ⥤ B} + [PreservesLimit (F ⋙ G) H] : PreservesLimit G H where + preserves {c} hc := by + refine ⟨isLimitWhiskerEquiv F _ ?_⟩ + let hc' := isLimitOfPreserves H ((isLimitWhiskerEquiv F _).symm hc) + exact IsLimit.ofIsoLimit hc' (Cones.ext (Iso.refl _) (by simp)) + +theorem reflectsLimit_of_comp {B : Type u₄} [Category.{v₄} B] {H : E ⥤ B} + [ReflectsLimit (F ⋙ G) H] : ReflectsLimit G H where + reflects {c} hc := by + refine ⟨isLimitWhiskerEquiv F _ (isLimitOfReflects H ?_)⟩ + let hc' := (isLimitWhiskerEquiv F _).symm hc + exact IsLimit.ofIsoLimit hc' (Cones.ext (Iso.refl _) (by simp)) + +/-- If `F` is initial and `F ⋙ G` creates limits of `H`, then so does `G`. -/ +def createsLimitOfComp {B : Type u₄} [Category.{v₄} B] {H : E ⥤ B} + [CreatesLimit (F ⋙ G) H] : CreatesLimit G H where + reflects := (reflectsLimit_of_comp F).reflects + lifts {c} hc := by + refine ⟨(extendCone (F := F)).obj (liftLimit ((isLimitWhiskerEquiv F _).symm hc)), ?_⟩ + let i := liftedLimitMapsToOriginal (K := (F ⋙ G)) ((isLimitWhiskerEquiv F _).symm hc) + refine ?_ ≪≫ ((extendCone (F := F)).mapIso i) ≪≫ ((conesEquiv F (G ⋙ H)).counitIso.app _) + exact Cones.ext (Iso.refl _) + include F in theorem hasLimitsOfShape_of_initial [HasLimitsOfShape C E] : HasLimitsOfShape D E where has_limit := fun _ => hasLimit_of_comp F +include F in +theorem preservesLimitsOfShape_of_initial {B : Type u₄} [Category.{v₄} B] (H : E ⥤ B) + [PreservesLimitsOfShape C H] : PreservesLimitsOfShape D H where + preservesLimit := preservesLimit_of_comp F + +include F in +theorem reflectsLimitsOfShape_of_initial {B : Type u₄} [Category.{v₄} B] (H : E ⥤ B) + [ReflectsLimitsOfShape C H] : ReflectsLimitsOfShape D H where + reflectsLimit := reflectsLimit_of_comp F + +include F in +/-- If `H` creates limits of shape `C` and `F : C ⥤ D` is initial, then `H` creates limits of shape +`D`. -/ +def createsLimitsOfShapeOfInitial {B : Type u₄} [Category.{v₄} B] (H : E ⥤ B) + [CreatesLimitsOfShape C H] : CreatesLimitsOfShape D H where + CreatesLimit := createsLimitOfComp F + end Initial section From dbcafd024242e06f4dbf2169991776331edca4a6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Nov 2024 15:32:00 +0000 Subject: [PATCH 292/829] feat: vertical line test for functions, group homs, linear maps (#18822) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Prove the vertical line test for monoid homomorphisms/isomorphisms. Let `f : G → H × I` be a homomorphism to a product of monoids. Assume that `f` is surjective on the first factor and that the image of `f` intersects every "vertical line" `{(h, i) | i : I}` at most once. Then the image of `f` is the graph of some monoid homomorphism `f' : H → I`. Furthermore, if `f` is also surjective on the second factor and its image intersects every "horizontal line" `{(h, i) | h : H}` at most once, then `f'` is actually an isomorphism `f' : H ≃ I`. Co-authored-by: David Loeffler --- Mathlib.lean | 1 + Mathlib/Algebra/Group/Graph.lean | 231 +++++++++++++++++++++++++++++++ Mathlib/Algebra/Group/Prod.lean | 2 + Mathlib/Data/Prod/Basic.lean | 5 +- Mathlib/Data/Set/Function.lean | 51 +++++++ Mathlib/LinearAlgebra/Prod.lean | 82 +++++++++++ 6 files changed, 371 insertions(+), 1 deletion(-) create mode 100644 Mathlib/Algebra/Group/Graph.lean diff --git a/Mathlib.lean b/Mathlib.lean index 02bb3ae68c288..979ce54ccf805 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -271,6 +271,7 @@ import Mathlib.Algebra.Group.Fin.Basic import Mathlib.Algebra.Group.Fin.Tuple import Mathlib.Algebra.Group.FiniteSupport import Mathlib.Algebra.Group.ForwardDiff +import Mathlib.Algebra.Group.Graph import Mathlib.Algebra.Group.Hom.Basic import Mathlib.Algebra.Group.Hom.CompTypeclasses import Mathlib.Algebra.Group.Hom.Defs diff --git a/Mathlib/Algebra/Group/Graph.lean b/Mathlib/Algebra/Group/Graph.lean new file mode 100644 index 0000000000000..a3cf4bee29867 --- /dev/null +++ b/Mathlib/Algebra/Group/Graph.lean @@ -0,0 +1,231 @@ +/- +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, David Loeffler +-/ +import Mathlib.Algebra.Group.Subgroup.Ker + +/-! +# Vertical line test for group homs + +This file proves the vertical line test for monoid homomorphisms/isomorphisms. + +Let `f : G → H × I` be a homomorphism to a product of monoids. Assume that `f` is surjective on the +first factor and that the image of `f` intersects every "vertical line" `{(h, i) | i : I}` at most +once. Then the image of `f` is the graph of some monoid homomorphism `f' : H → I`. + +Furthermore, if `f` is also surjective on the second factor and its image intersects every +"horizontal line" `{(h, i) | h : H}` at most once, then `f'` is actually an isomorphism +`f' : H ≃ I`. + +We also prove specialised versions when `f` is the inclusion of a subgroup of the direct product. +(The version for general homomorphisms can easily be reduced to this special case, but the +homomorphism version is more flexible in applications.) +-/ + +open Function Set + +variable {G H I : Type*} + +section Monoid +variable [Monoid G] [Monoid H] [Monoid I] + +namespace MonoidHom + +/-- The graph of a monoid homomorphism as a submonoid. + +See also `MonoidHom.graph` for the graph as a subgroup. -/ +@[to_additive +"The graph of a monoid homomorphism as a submonoid. + +See also `AddMonoidHom.graph` for the graph as a subgroup."] +def mgraph (f : G →* H) : Submonoid (G × H) where + carrier := {x | f x.1 = x.2} + one_mem' := map_one f + mul_mem' {x y} := by simp +contextual + +-- TODO: Can `to_additive` be smarter about `simps`? +attribute [simps! coe] mgraph +attribute [simps! coe] AddMonoidHom.mgraph +set_option linter.existingAttributeWarning false in +attribute [to_additive existing] coe_mgraph + +@[to_additive (attr := simp)] +lemma mem_mgraph {f : G →* H} {x : G × H} : x ∈ f.mgraph ↔ f x.1 = x.2 := .rfl + +@[to_additive] +lemma mgraph_eq_mrange_prod (f : G →* H) : f.mgraph = mrange ((id _).prod f) := by aesop + +/-- **Vertical line test** for monoid homomorphisms. + +Let `f : G → H × I` be a homomorphism to a product of monoids. Assume that `f` is surjective on the +first factor and that the image of `f` intersects every "vertical line" `{(h, i) | i : I}` at most +once. Then the image of `f` is the graph of some monoid homomorphism `f' : H → I`. -/ +@[to_additive "**Vertical line test** for monoid homomorphisms. + +Let `f : G → H × I` be a homomorphism to a product of monoids. Assume that `f` is surjective on the +first factor and that the image of `f` intersects every \"vertical line\" `{(h, i) | i : I}` at most +once. Then the image of `f` is the graph of some monoid homomorphism `f' : H → I`."] +lemma exists_mrange_eq_mgraph {f : G →* H × I} (hf₁ : Surjective (Prod.fst ∘ f)) + (hf : ∀ g₁ g₂, (f g₁).1 = (f g₂).1 → (f g₁).2 = (f g₂).2) : + ∃ f' : H →* I, mrange f = f'.mgraph := by + obtain ⟨f', hf'⟩ := exists_range_eq_graphOn_univ hf₁ hf + simp only [Set.ext_iff, Set.mem_range, mem_graphOn, mem_univ, true_and, Prod.forall] at hf' + use + { toFun := f' + map_one' := (hf' _ _).1 ⟨1, map_one _⟩ + map_mul' := by + simp_rw [hf₁.forall] + rintro g₁ g₂ + exact (hf' _ _).1 ⟨g₁ * g₂, by simp [Prod.ext_iff, (hf' (f _).1 _).1 ⟨_, rfl⟩]⟩ } + simpa [SetLike.ext_iff] using hf' + +/-- **Line test** for monoid isomorphisms. + +Let `f : G → H × I` be a homomorphism to a product of monoids. Assume that `f` is surjective on both +factors and that the image of `f` intersects every "vertical line" `{(h, i) | i : I}` and every +"horizontal line" `{(h, i) | h : H}` at most once. Then the image of `f` is the graph of some monoid +isomorphism `f' : H ≃ I`. -/ +@[to_additive "**Line test** for monoid isomorphisms. + +Let `f : G → H × I` be a homomorphism to a product of monoids. Assume that `f` is surjective on both +factors and that the image of `f` intersects every \"vertical line\" `{(h, i) | i : I}` and every +\"horizontal line\" `{(h, i) | h : H}` at most once. Then the image of `f` is the graph of some +monoid isomorphism `f' : H ≃ I`."] +lemma exists_mulEquiv_mrange_eq_mgraph {f : G →* H × I} (hf₁ : Surjective (Prod.fst ∘ f)) + (hf₂ : Surjective (Prod.snd ∘ f)) (hf : ∀ g₁ g₂, (f g₁).1 = (f g₂).1 ↔ (f g₁).2 = (f g₂).2) : + ∃ e : H ≃* I, mrange f = e.toMonoidHom.mgraph := by + obtain ⟨e₁, he₁⟩ := f.exists_mrange_eq_mgraph hf₁ fun _ _ ↦ (hf _ _).1 + obtain ⟨e₂, he₂⟩ := (MulEquiv.prodComm.toMonoidHom.comp f).exists_mrange_eq_mgraph (by simpa) <| + by simp [hf] + have he₁₂ h i : e₁ h = i ↔ e₂ i = h := by + rw [SetLike.ext_iff] at he₁ he₂ + aesop (add simp [Prod.swap_eq_iff_eq_swap]) + exact ⟨ + { toFun := e₁ + map_mul' := e₁.map_mul' + invFun := e₂ + left_inv := fun h ↦ by rw [← he₁₂] + right_inv := fun i ↦ by rw [he₁₂] }, he₁⟩ + +end MonoidHom + +/-- **Vertical line test** for monoid homomorphisms. + +Let `G ≤ H × I` be a submonoid of a product of monoids. Assume that `G` maps bijectively to the +first factor. Then `G` is the graph of some monoid homomorphism `f : H → I`. -/ +@[to_additive "**Vertical line test** for additive monoid homomorphisms. + +Let `G ≤ H × I` be a submonoid of a product of monoids. Assume that `G` surjects onto the first +factor and `G` intersects every \"vertical line\" `{(h, i) | i : I}` at most once. Then `G` is the +graph of some monoid homomorphism `f : H → I`."] +lemma Submonoid.exists_eq_mgraph {G : Submonoid (H × I)} (hG₁ : Bijective (Prod.fst ∘ G.subtype)) : + ∃ f : H →* I, G = f.mgraph := by + simpa using MonoidHom.exists_mrange_eq_mgraph hG₁.surjective + fun a b h ↦ congr_arg (Prod.snd ∘ G.subtype) (hG₁.injective h) + +/-- **Goursat's lemma** for monoid isomorphisms. + +Let `G ≤ H × I` be a submonoid of a product of monoids. Assume that the natural maps from `G` to +both factors are bijective. Then `G` is the graph of some isomorphism `f : H ≃* I`. -/ +@[to_additive "**Goursat's lemma** for additive monoid isomorphisms. + +Let `G ≤ H × I` be a submonoid of a product of additive monoids. Assume that the natural maps from +`G` to both factors are bijective. Then `G` is the graph of some isomorphism `f : H ≃+ I`."] +lemma Submonoid.exists_mulEquiv_eq_mgraph {G : Submonoid (H × I)} + (hG₁ : Bijective (Prod.fst ∘ G.subtype)) (hG₂ : Bijective (Prod.snd ∘ G.subtype)) : + ∃ e : H ≃* I, G = e.toMonoidHom.mgraph := by + simpa using MonoidHom.exists_mulEquiv_mrange_eq_mgraph hG₁.surjective hG₂.surjective + fun _ _ ↦ hG₁.injective.eq_iff.trans hG₂.injective.eq_iff.symm + +end Monoid + +section Group +variable [Group G] [Group H] [Group I] + +namespace MonoidHom + +/-- The graph of a group homomorphism as a subgroup. + +See also `MonoidHom.mgraph` for the graph as a submonoid. -/ +@[to_additive +"The graph of a group homomorphism as a subgroup. + +See also `AddMonoidHom.mgraph` for the graph as a submonoid."] +def graph (f : G →* H) : Subgroup (G × H) where + toSubmonoid := f.mgraph + inv_mem' {x} := by simp +contextual + +-- TODO: Can `to_additive` be smarter about `simps`? +attribute [simps! coe toSubmonoid] graph +attribute [simps! coe toAddSubmonoid] AddMonoidHom.graph +set_option linter.existingAttributeWarning false in +attribute [to_additive existing] coe_graph graph_toSubmonoid + +@[to_additive] +lemma mem_graph {f : G →* H} {x : G × H} : x ∈ f.mgraph ↔ f x.1 = x.2 := .rfl + +@[to_additive] +lemma graph_eq_range_prod (f : G →* H) : f.graph = range ((id _).prod f) := by aesop + +/-- **Vertical line test** for group homomorphisms. + +Let `f : G → H × I` be a homomorphism to a product of groups. Assume that `f` is surjective on the +first factor and that the image of `f` intersects every "vertical line" `{(h, i) | i : I}` at most +once. Then the image of `f` is the graph of some group homomorphism `f' : H → I`. -/ +@[to_additive "**Vertical line test** for group homomorphisms. + +Let `f : G → H × I` be a homomorphism to a product of groups. Assume that `f` is surjective on the +first factor and that the image of `f` intersects every \"vertical line\" `{(h, i) | i : I}` at most +once. Then the image of `f` is the graph of some group homomorphism `f' : H → I`."] +lemma exists_range_eq_graph {f : G →* H × I} (hf₁ : Surjective (Prod.fst ∘ f)) + (hf : ∀ g₁ g₂, (f g₁).1 = (f g₂).1 → (f g₁).2 = (f g₂).2) : + ∃ f' : H →* I, range f = f'.graph := by + simpa [SetLike.ext_iff] using exists_mrange_eq_mgraph hf₁ hf + +/-- **Line test** for group isomorphisms. + +Let `f : G → H × I` be a homomorphism to a product of groups. Assume that `f` is surjective on both +factors and that the image of `f` intersects every "vertical line" `{(h, i) | i : I}` and every +"horizontal line" `{(h, i) | h : H}` at most once. Then the image of `f` is the graph of some group +isomorphism `f' : H ≃ I`. -/ +@[to_additive "**Line test** for monoid isomorphisms. + +Let `f : G → H × I` be a homomorphism to a product of groups. Assume that `f` is surjective on both +factors and that the image of `f` intersects every \"vertical line\" `{(h, i) | i : I}` and every +\"horizontal line\" `{(h, i) | h : H}` at most once. Then the image of `f` is the graph of some +group isomorphism `f' : H ≃ I`."] +lemma exists_mulEquiv_range_eq_graph {f : G →* H × I} (hf₁ : Surjective (Prod.fst ∘ f)) + (hf₂ : Surjective (Prod.snd ∘ f)) (hf : ∀ g₁ g₂, (f g₁).1 = (f g₂).1 ↔ (f g₁).2 = (f g₂).2) : + ∃ e : H ≃* I, range f = e.toMonoidHom.graph := by + simpa [SetLike.ext_iff] using exists_mulEquiv_mrange_eq_mgraph hf₁ hf₂ hf + +end MonoidHom + +/-- **Vertical line test** for group homomorphisms. + +Let `G ≤ H × I` be a subgroup of a product of monoids. Assume that `G` maps bijectively to the +first factor. Then `G` is the graph of some monoid homomorphism `f : H → I`. -/ +@[to_additive "**Vertical line test** for additive monoid homomorphisms. + +Let `G ≤ H × I` be a submonoid of a product of monoids. Assume that `G` surjects onto the first +factor and `G` intersects every \"vertical line\" `{(h, i) | i : I}` at most once. Then `G` is the +graph of some monoid homomorphism `f : H → I`."] +lemma Subgroup.exists_eq_graph {G : Subgroup (H × I)} (hG₁ : Bijective (Prod.fst ∘ G.subtype)) : + ∃ f : H →* I, G = f.graph := by + simpa [SetLike.ext_iff] using Submonoid.exists_eq_mgraph hG₁ + +/-- **Goursat's lemma** for monoid isomorphisms. + +Let `G ≤ H × I` be a submonoid of a product of monoids. Assume that the natural maps from `G` to +both factors are bijective. Then `G` is the graph of some isomorphism `f : H ≃* I`. -/ +@[to_additive "**Goursat's lemma** for additive monoid isomorphisms. + +Let `G ≤ H × I` be a submonoid of a product of additive monoids. Assume that the natural maps from +`G` to both factors are bijective. Then `G` is the graph of some isomorphism `f : H ≃+ I`."] +lemma Subgroup.exists_mulEquiv_eq_graph {G : Subgroup (H × I)} + (hG₁ : Bijective (Prod.fst ∘ G.subtype)) (hG₂ : Bijective (Prod.snd ∘ G.subtype)) : + ∃ e : H ≃* I, G = e.toMonoidHom.graph := by + simpa [SetLike.ext_iff] using Submonoid.exists_mulEquiv_eq_mgraph hG₁ hG₂ + +end Group diff --git a/Mathlib/Algebra/Group/Prod.lean b/Mathlib/Algebra/Group/Prod.lean index 02781d10d0067..e6e9519ae6483 100644 --- a/Mathlib/Algebra/Group/Prod.lean +++ b/Mathlib/Algebra/Group/Prod.lean @@ -149,6 +149,8 @@ theorem mk_div_mk [Div G] [Div H] (x₁ x₂ : G) (y₁ y₂ : H) : theorem swap_div [Div G] [Div H] (a b : G × H) : (a / b).swap = a.swap / b.swap := rfl +@[to_additive] lemma div_def [Div M] [Div N] (a b : M × N) : a / b = (a.1 / b.1, a.2 / b.2) := rfl + @[to_additive] instance instSemigroup [Semigroup M] [Semigroup N] : Semigroup (M × N) := { mul_assoc := fun _ _ _ => mk.inj_iff.mpr ⟨mul_assoc _ _ _, mul_assoc _ _ _⟩ } diff --git a/Mathlib/Data/Prod/Basic.lean b/Mathlib/Data/Prod/Basic.lean index f42386ead9042..98f729aaec3e6 100644 --- a/Mathlib/Data/Prod/Basic.lean +++ b/Mathlib/Data/Prod/Basic.lean @@ -5,12 +5,13 @@ Authors: Johannes Hölzl -/ import Mathlib.Logic.Function.Defs import Mathlib.Logic.Function.Iterate +import Aesop import Mathlib.Tactic.Inhabit /-! # Extra facts about `Prod` -This file defines `Prod.swap : α × β → β × α` and proves various simple lemmas about `Prod`. +This file proves various simple lemmas about `Prod`. It also defines better delaborators for product projections. -/ @@ -20,6 +21,8 @@ variable {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} namespace Prod +lemma swap_eq_iff_eq_swap {x : α × β} {y : β × α} : x.swap = y ↔ x = y.swap := by aesop + def mk.injArrow {x₁ : α} {y₁ : β} {x₂ : α} {y₂ : β} : (x₁, y₁) = (x₂, y₂) → ∀ ⦃P : Sort*⦄, (x₁ = x₂ → y₁ = y₂ → P) → P := fun h₁ _ h₂ ↦ Prod.noConfusion h₁ h₂ diff --git a/Mathlib/Data/Set/Function.lean b/Mathlib/Data/Set/Function.lean index 0a7bacfba799b..77dbaaa172001 100644 --- a/Mathlib/Data/Set/Function.lean +++ b/Mathlib/Data/Set/Function.lean @@ -1665,4 +1665,55 @@ lemma bijOn_swap (ha : a ∈ s) (hb : b ∈ s) : BijOn (swap a b) s s := end Equiv +/-! ### Vertical line test -/ + +namespace Set + +/-- **Vertical line test** for functions. + +Let `f : α → β × γ` be a function to a product. Assume that `f` is surjective on the first factor +and that the image of `f` intersects every "vertical line" `{(b, c) | c : γ}` at most once. +Then the image of `f` is the graph of some monoid homomorphism `f' : β → γ`. -/ +lemma exists_range_eq_graphOn_univ {f : α → β × γ} (hf₁ : Surjective (Prod.fst ∘ f)) + (hf : ∀ g₁ g₂, (f g₁).1 = (f g₂).1 → (f g₁).2 = (f g₂).2) : + ∃ f' : β → γ, range f = univ.graphOn f' := by + refine ⟨fun h ↦ (f (hf₁ h).choose).snd, ?_⟩ + ext x + simp only [mem_range, comp_apply, mem_graphOn, mem_univ, true_and] + refine ⟨?_, fun hi ↦ ⟨(hf₁ x.1).choose, Prod.ext (hf₁ x.1).choose_spec hi⟩⟩ + rintro ⟨g, rfl⟩ + exact hf _ _ (hf₁ (f g).1).choose_spec + +/-- **Line test** for equivalences. + +Let `f : α → β × γ` be a homomorphism to a product of monoids. Assume that `f` is surjective on both +factors and that the image of `f` intersects every "vertical line" `{(b, c) | c : γ}` and every +"horizontal line" `{(b, c) | b : β}` at most once. Then the image of `f` is the graph of some +equivalence `f' : β ≃ γ`. -/ +lemma exists_equiv_range_eq_graphOn_univ {f : α → β × γ} (hf₁ : Surjective (Prod.fst ∘ f)) + (hf₂ : Surjective (Prod.snd ∘ f)) (hf : ∀ g₁ g₂, (f g₁).1 = (f g₂).1 ↔ (f g₁).2 = (f g₂).2) : + ∃ e : β ≃ γ, range f = univ.graphOn e := by + obtain ⟨e₁, he₁⟩ := exists_range_eq_graphOn_univ hf₁ fun _ _ ↦ (hf _ _).1 + obtain ⟨e₂, he₂⟩ := exists_range_eq_graphOn_univ (f := Equiv.prodComm _ _ ∘ f) (by simpa) <| + by simp [hf] + have he₁₂ h i : e₁ h = i ↔ e₂ i = h := by + rw [Set.ext_iff] at he₁ he₂ + aesop (add simp [Prod.swap_eq_iff_eq_swap]) + exact ⟨ + { toFun := e₁ + invFun := e₂ + left_inv := fun h ↦ by rw [← he₁₂] + right_inv := fun i ↦ by rw [he₁₂] }, he₁⟩ + +/-- **Vertical line test** for functions. + +Let `s : Set (β × γ)` be a set in a product. Assume that `s` maps bijectively to the first factor. +Then `s` is the graph of some function `f : β → γ`. -/ +lemma exists_eq_mgraphOn_univ {s : Set (β × γ)} + (hs₁ : Bijective (Prod.fst ∘ (Subtype.val : s → β × γ))) : ∃ f : β → γ, s = univ.graphOn f := by + simpa using exists_range_eq_graphOn_univ hs₁.surjective + fun a b h ↦ congr_arg (Prod.snd ∘ (Subtype.val : s → β × γ)) (hs₁.injective h) + +end Set + set_option linter.style.longFile 1800 diff --git a/Mathlib/LinearAlgebra/Prod.lean b/Mathlib/LinearAlgebra/Prod.lean index 7bbb4ac2f4b65..23732c0a35284 100644 --- a/Mathlib/LinearAlgebra/Prod.lean +++ b/Mathlib/LinearAlgebra/Prod.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Kevin Buzzard, Yury Kudryashov, Eric Wieser -/ import Mathlib.Algebra.Algebra.Prod +import Mathlib.Algebra.Group.Graph import Mathlib.LinearAlgebra.Span.Basic import Mathlib.Order.PartialSups @@ -965,3 +966,84 @@ theorem graph_eq_range_prod : f.graph = range (LinearMap.id.prod f) := by end Graph end LinearMap + +section LineTest + +open Set Function + +variable {R S G H I : Type*} + [Semiring R] [Semiring S] {σ : R →+* S} [RingHomSurjective σ] + [AddCommMonoid G] [Module R G] + [AddCommMonoid H] [Module S H] + [AddCommMonoid I] [Module S I] + +/-- **Vertical line test** for module homomorphisms. + +Let `f : G → H × I` be a linear (or semilinear) map to a product. Assume that `f` is surjective on +the first factor and that the image of `f` intersects every "vertical line" `{(h, i) | i : I}` at +most once. Then the image of `f` is the graph of some linear map `f' : H → I`. -/ +lemma LinearMap.exists_range_eq_graph {f : G →ₛₗ[σ] H × I} (hf₁ : Surjective (Prod.fst ∘ f)) + (hf : ∀ g₁ g₂, (f g₁).1 = (f g₂).1 → (f g₁).2 = (f g₂).2) : + ∃ f' : H →ₗ[S] I, LinearMap.range f = LinearMap.graph f' := by + obtain ⟨f', hf'⟩ := AddMonoidHom.exists_mrange_eq_mgraph (I := I) (f := f) hf₁ hf + simp only [SetLike.ext_iff, AddMonoidHom.mem_mrange, AddMonoidHom.coe_coe, + AddMonoidHom.mem_mgraph] at hf' + use + { toFun := f'.toFun + map_add' := f'.map_add' + map_smul' := by + intro s h + simp only [ZeroHom.toFun_eq_coe, AddMonoidHom.toZeroHom_coe, RingHom.id_apply] + refine (hf' (s • h, _)).mp ?_ + rw [← Prod.smul_mk, ← LinearMap.mem_range] + apply Submodule.smul_mem + rw [LinearMap.mem_range, hf'] } + ext x + simpa only [mem_range, Eq.comm, ZeroHom.toFun_eq_coe, AddMonoidHom.toZeroHom_coe, mem_graph_iff, + coe_mk, AddHom.coe_mk, AddMonoidHom.coe_coe, Set.mem_range] using hf' x + +/-- **Vertical line test** for module homomorphisms. + +Let `G ≤ H × I` be a submodule of a product of modules. Assume that `G` maps bijectively to the +first factor. Then `G` is the graph of some module homomorphism `f : H →ₗ[R] I`. -/ +lemma Submodule.exists_eq_graph {G : Submodule S (H × I)} (hf₁ : Bijective (Prod.fst ∘ G.subtype)) : + ∃ f : H →ₗ[S] I, G = LinearMap.graph f := by + simpa only [range_subtype] using LinearMap.exists_range_eq_graph hf₁.surjective + (fun a b h ↦ congr_arg (Prod.snd ∘ G.subtype) (hf₁.injective h)) + +/-- **Line test** for module isomorphisms. + +Let `f : G → H × I` be a homomorphism to a product of modules. Assume that `f` is surjective onto +both factors and that the image of `f` intersects every "vertical line" `{(h, i) | i : I}` and every +"horizontal line" `{(h, i) | h : H}` at most once. Then the image of `f` is the graph of some +module isomorphism `f' : H ≃ I`. -/ +lemma LinearMap.exists_linearEquiv_eq_graph {f : G →ₛₗ[σ] H × I} (hf₁ : Surjective (Prod.fst ∘ f)) + (hf₂ : Surjective (Prod.snd ∘ f)) (hf : ∀ g₁ g₂, (f g₁).1 = (f g₂).1 ↔ (f g₁).2 = (f g₂).2) : + ∃ e : H ≃ₗ[S] I, range f = e.toLinearMap.graph := by + obtain ⟨e₁, he₁⟩ := f.exists_range_eq_graph hf₁ fun _ _ ↦ (hf _ _).1 + obtain ⟨e₂, he₂⟩ := ((LinearEquiv.prodComm _ _ _).toLinearMap.comp f).exists_range_eq_graph + (by simpa) <| by simp [hf] + have he₁₂ h i : e₁ h = i ↔ e₂ i = h := by + simp only [SetLike.ext_iff, LinearMap.mem_graph_iff] at he₁ he₂ + rw [Eq.comm, ← he₁ (h, i), Eq.comm, ← he₂ (i, h)] + simp only [mem_range, coe_comp, LinearEquiv.coe_coe, Function.comp_apply, + LinearEquiv.prodComm_apply, Prod.swap_eq_iff_eq_swap, Prod.swap_prod_mk] + exact ⟨ + { toFun := e₁ + map_smul' := e₁.map_smul' + map_add' := e₁.map_add' + invFun := e₂ + left_inv := fun h ↦ by rw [← he₁₂] + right_inv := fun i ↦ by rw [he₁₂] }, he₁⟩ + +/-- **Goursat's lemma** for module isomorphisms. + +Let `G ≤ H × I` be a submodule of a product of modules. Assume that the natural maps from `G` to +both factors are bijective. Then `G` is the graph of some module isomorphism `f : H ≃ I`. -/ +lemma Submodule.exists_equiv_eq_graph {G : Submodule S (H × I)} + (hG₁ : Bijective (Prod.fst ∘ G.subtype)) (hG₂ : Bijective (Prod.snd ∘ G.subtype)) : + ∃ e : H ≃ₗ[S] I, G = e.toLinearMap.graph := by + simpa only [range_subtype] using LinearMap.exists_linearEquiv_eq_graph + hG₁.surjective hG₂.surjective fun _ _ ↦ hG₁.injective.eq_iff.trans hG₂.injective.eq_iff.symm + +end LineTest From 04cd3d0afc11a0bcb07934a8e2f12a72793b2f28 Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 25 Nov 2024 15:32:01 +0000 Subject: [PATCH 293/829] fix: remove `ref: master` in zulip emoji reaction script (#19465) This way, the script will extract the python file from the current PR, allowing testing and debugging of changes in its behaviour. --- .github/workflows/maintainer_bors.yml | 1 - 1 file changed, 1 deletion(-) diff --git a/.github/workflows/maintainer_bors.yml b/.github/workflows/maintainer_bors.yml index 67e3596b6a67e..10a12e5f3a439 100644 --- a/.github/workflows/maintainer_bors.yml +++ b/.github/workflows/maintainer_bors.yml @@ -103,7 +103,6 @@ jobs: - uses: actions/checkout@v4 with: - ref: master sparse-checkout: | scripts/zulip_emoji_merge_delegate.py From 15ea3626fdf7430332c8b5a605e53498a0094512 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 25 Nov 2024 16:00:55 +0000 Subject: [PATCH 294/829] chore(Topology/../LiminfLimsup): change assumptions of 2 instances (#18919) - Assume `LinearOrder` instead of `Preorder`. - Assume `ClosedIciTopology`/`ClosedIicTopology` instead of `OrderTopology`. `OrderTopology` is usually not the right topology for non-linear orders, so restricting from `Preorder` to `LinearOrder` doesn't break anything. --- .../Topology/Algebra/Order/LiminfLimsup.lean | 20 +++++++++---------- 1 file changed, 9 insertions(+), 11 deletions(-) diff --git a/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean b/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean index bc4753090a8e3..9f2b202d56789 100644 --- a/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean +++ b/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean @@ -130,20 +130,18 @@ instance (priority := 100) OrderTop.to_BoundedLENhdsClass [OrderTop α] : Bounde instance (priority := 100) OrderBot.to_BoundedGENhdsClass [OrderBot α] : BoundedGENhdsClass α := ⟨fun _a ↦ isBounded_ge_of_bot⟩ --- See note [lower instance priority] -instance (priority := 100) OrderTopology.to_BoundedLENhdsClass [IsDirected α (· ≤ ·)] - [OrderTopology α] : BoundedLENhdsClass α := - ⟨fun a ↦ - ((isTop_or_exists_gt a).elim fun h ↦ ⟨a, Eventually.of_forall h⟩) <| - Exists.imp fun _b ↦ ge_mem_nhds⟩ +end Preorder -- See note [lower instance priority] -instance (priority := 100) OrderTopology.to_BoundedGENhdsClass [IsDirected α (· ≥ ·)] - [OrderTopology α] : BoundedGENhdsClass α := - ⟨fun a ↦ ((isBot_or_exists_lt a).elim fun h ↦ ⟨a, Eventually.of_forall h⟩) <| - Exists.imp fun _b ↦ le_mem_nhds⟩ +instance (priority := 100) BoundedLENhdsClass.of_closedIciTopology [LinearOrder α] + [TopologicalSpace α] [ClosedIciTopology α] : BoundedLENhdsClass α := + ⟨fun a ↦ ((isTop_or_exists_gt a).elim fun h ↦ ⟨a, Eventually.of_forall h⟩) <| + Exists.imp fun _b ↦ eventually_le_nhds⟩ -end Preorder +-- See note [lower instance priority] +instance (priority := 100) BoundedGENhdsClass.of_closedIicTopology [LinearOrder α] + [TopologicalSpace α] [ClosedIicTopology α] : BoundedGENhdsClass α := + inferInstanceAs <| BoundedGENhdsClass αᵒᵈᵒᵈ section LiminfLimsup From 435fc2bcedc84fd97664f7c4ece07d979765063c Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Mon, 25 Nov 2024 16:18:55 +0000 Subject: [PATCH 295/829] feat(AlgebraicGeometry): covers of schemes over a base (#19096) We introduce a class on covers containing the data of structure morphisms over a base and provide basic constructions. Co-authored-by: Christian Merten <136261474+chrisflav@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/AlgebraicGeometry/Cover/Over.lean | 183 ++++++++++++++++++ .../Morphisms/UnderlyingMap.lean | 12 ++ Mathlib/AlgebraicGeometry/Over.lean | 7 + Mathlib/CategoryTheory/Comma/OverClass.lean | 15 ++ .../Limits/MorphismProperty.lean | 2 +- .../MorphismProperty/Comma.lean | 8 + 7 files changed, 227 insertions(+), 1 deletion(-) create mode 100644 Mathlib/AlgebraicGeometry/Cover/Over.lean diff --git a/Mathlib.lean b/Mathlib.lean index 979ce54ccf805..c6a763cc2c95e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -931,6 +931,7 @@ import Mathlib.AlgebraicGeometry.AffineScheme import Mathlib.AlgebraicGeometry.AffineSpace import Mathlib.AlgebraicGeometry.Cover.MorphismProperty import Mathlib.AlgebraicGeometry.Cover.Open +import Mathlib.AlgebraicGeometry.Cover.Over import Mathlib.AlgebraicGeometry.EllipticCurve.Affine import Mathlib.AlgebraicGeometry.EllipticCurve.DivisionPolynomial.Basic import Mathlib.AlgebraicGeometry.EllipticCurve.DivisionPolynomial.Degree diff --git a/Mathlib/AlgebraicGeometry/Cover/Over.lean b/Mathlib/AlgebraicGeometry/Cover/Over.lean new file mode 100644 index 0000000000000..e6b2a8b4e0e23 --- /dev/null +++ b/Mathlib/AlgebraicGeometry/Cover/Over.lean @@ -0,0 +1,183 @@ +/- +Copyright (c) 2024 Christian Merten. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Christian Merten +-/ +import Mathlib.AlgebraicGeometry.Morphisms.UnderlyingMap +import Mathlib.CategoryTheory.Limits.MorphismProperty + +/-! + +# Covers of schemes over a base + +In this file we define the typeclass `Cover.Over`. For a cover `𝒰` of an `S`-scheme `X`, +the datum `𝒰.Over S` contains `S`-scheme structures on the components of `𝒰` and asserts +that the component maps are morphisms of `S`-schemes. + +We provide instances of `𝒰.Over S` for standard constructions on covers. + +-/ + +universe v u + +noncomputable section + +open CategoryTheory Limits + +namespace AlgebraicGeometry.Scheme + +variable {P : MorphismProperty Scheme.{u}} (S : Scheme.{u}) + +/-- Bundle an `S`-scheme with `P` into an object of `P.Over ⊤ S`. -/ +abbrev asOverProp (X : Scheme.{u}) (S : Scheme.{u}) [X.Over S] (h : P (X ↘ S)) : P.Over ⊤ S := + ⟨X.asOver S, h⟩ + +/-- Bundle an `S`-morphism of `S`-scheme with `P` into a morphism in `P.Over ⊤ S`. -/ +abbrev Hom.asOverProp {X Y : Scheme.{u}} (f : X.Hom Y) (S : Scheme.{u}) [X.Over S] [Y.Over S] + [f.IsOver S] {hX : P (X ↘ S)} {hY : P (Y ↘ S)} : X.asOverProp S hX ⟶ Y.asOverProp S hY := + ⟨f.asOver S, trivial, trivial⟩ + +/-- A `P`-cover of a scheme `X` over `S` is a cover, where the components are over `S` and the +component maps commute with the structure morphisms. -/ +protected class Cover.Over {P : MorphismProperty Scheme.{u}} {X : Scheme.{u}} [X.Over S] + (𝒰 : X.Cover P) where + over (j : 𝒰.J) : (𝒰.obj j).Over S := by infer_instance + isOver_map (j : 𝒰.J) : (𝒰.map j).IsOver S := by infer_instance + +attribute [instance] Cover.Over.over Cover.Over.isOver_map + +instance [P.ContainsIdentities] [P.RespectsIso] {X Y : Scheme.{u}} (f : X ⟶ Y) [X.Over S] [Y.Over S] + [f.IsOver S] [IsIso f] : (coverOfIsIso (P := P) f).Over S where + over _ := inferInstanceAs <| X.Over S + isOver_map _ := inferInstanceAs <| f.IsOver S + +section + +variable [P.IsStableUnderBaseChange] [IsJointlySurjectivePreserving P] +variable {X W : Scheme.{u}} (𝒰 : X.Cover P) (f : W ⟶ X) [W.Over S] [X.Over S] + [𝒰.Over S] [f.IsOver S] + +/-- The pullback of a cover of `S`-schemes along a morphism of `S`-schemes. This is not +definitionally equal to `AlgebraicGeometry.Scheme.Cover.pullbackCover`, as here we take +the pullback in `Over S`, whose underlying scheme is only isomorphic but not equal to the +pullback in `Scheme`. -/ +@[simps] +def Cover.pullbackCoverOver : W.Cover P where + J := 𝒰.J + obj x := (pullback (f.asOver S) ((𝒰.map x).asOver S)).left + map x := (pullback.fst (f.asOver S) ((𝒰.map x).asOver S)).left + f x := 𝒰.f (f.base x) + covers x := (mem_range_iff_of_surjective ((𝒰.pullbackCover f).map (𝒰.f (f.base x))) _ + ((PreservesPullback.iso (Over.forget S) (f.asOver S) ((𝒰.map _).asOver S)).inv) + (PreservesPullback.iso_inv_fst _ _ _) x).mp ((𝒰.pullbackCover f).covers x) + map_prop j := by + dsimp only + rw [← Over.forget_map, ← PreservesPullback.iso_hom_fst, P.cancel_left_of_respectsIso] + exact P.pullback_fst _ _ (𝒰.map_prop j) + +instance (j : 𝒰.J) : ((𝒰.pullbackCoverOver S f).obj j).Over S where + hom := (pullback (f.asOver S) ((𝒰.map j).asOver S)).hom + +instance : (𝒰.pullbackCoverOver S f).Over S where + isOver_map j := { comp_over := by exact Over.w (pullback.fst (f.asOver S) ((𝒰.map j).asOver S)) } + +/-- A variant of `AlgebraicGeometry.Scheme.Cover.pullbackCoverOver` with the arguments in the +fiber products flipped. -/ +@[simps] +def Cover.pullbackCoverOver' : W.Cover P where + J := 𝒰.J + obj x := (pullback ((𝒰.map x).asOver S) (f.asOver S)).left + map x := (pullback.snd ((𝒰.map x).asOver S) (f.asOver S)).left + f x := 𝒰.f (f.base x) + covers x := (mem_range_iff_of_surjective ((𝒰.pullbackCover' f).map (𝒰.f (f.base x))) _ + ((PreservesPullback.iso (Over.forget S) ((𝒰.map _).asOver S) (f.asOver S)).inv) + (PreservesPullback.iso_inv_snd _ _ _) x).mp ((𝒰.pullbackCover' f).covers x) + map_prop j := by + dsimp only + rw [← Over.forget_map, ← PreservesPullback.iso_hom_snd, P.cancel_left_of_respectsIso] + exact P.pullback_snd _ _ (𝒰.map_prop j) + +instance (j : 𝒰.J) : ((𝒰.pullbackCoverOver' S f).obj j).Over S where + hom := (pullback ((𝒰.map j).asOver S) (f.asOver S)).hom + +instance : (𝒰.pullbackCoverOver' S f).Over S where + isOver_map j := { comp_over := by exact Over.w (pullback.snd ((𝒰.map j).asOver S) (f.asOver S)) } + +variable {Q : MorphismProperty Scheme.{u}} [Q.HasOfPostcompProperty Q] + [Q.IsStableUnderBaseChange] [Q.IsStableUnderComposition] + +variable (hX : Q (X ↘ S)) (hW : Q (W ↘ S)) (hQ : ∀ j, Q (𝒰.obj j ↘ S)) + +/-- The pullback of a cover of `S`-schemes with `Q` along a morphism of `S`-schemes. This is not +definitionally equal to `AlgebraicGeometry.Scheme.Cover.pullbackCover`, as here we take +the pullback in `Q.Over ⊤ S`, whose underlying scheme is only isomorphic but not equal to the +pullback in `Scheme`. -/ +@[simps (config := .lemmasOnly)] +def Cover.pullbackCoverOverProp : W.Cover P where + J := 𝒰.J + obj x := (pullback (f.asOverProp (hX := hW) (hY := hX) S) + ((𝒰.map x).asOverProp (hX := hQ x) (hY := hX) S)).left + map x := (pullback.fst (f.asOverProp S) ((𝒰.map x).asOverProp S)).left + f x := 𝒰.f (f.base x) + covers x := (mem_range_iff_of_surjective ((𝒰.pullbackCover f).map (𝒰.f (f.base x))) _ + ((PreservesPullback.iso (MorphismProperty.Over.forget Q _ _ ⋙ Over.forget S) + (f.asOverProp S) ((𝒰.map _).asOverProp S)).inv) + (PreservesPullback.iso_inv_fst _ _ _) x).mp ((𝒰.pullbackCover f).covers x) + map_prop j := by + dsimp only + rw [← Over.forget_map, MorphismProperty.Comma.toCommaMorphism_eq_hom, + ← MorphismProperty.Comma.forget_map, ← Functor.comp_map] + rw [← PreservesPullback.iso_hom_fst, P.cancel_left_of_respectsIso] + exact P.pullback_fst _ _ (𝒰.map_prop j) + +instance (j : 𝒰.J) : ((𝒰.pullbackCoverOverProp S f hX hW hQ).obj j).Over S where + hom := (pullback (f.asOverProp (hX := hW) (hY := hX) S) + ((𝒰.map j).asOverProp (hX := hQ j) (hY := hX) S)).hom + +instance : (𝒰.pullbackCoverOverProp S f hX hW hQ).Over S where + isOver_map j := + { comp_over := by exact (pullback.fst (f.asOverProp S) ((𝒰.map j).asOverProp S)).w } + +/-- A variant of `AlgebraicGeometry.Scheme.Cover.pullbackCoverOverProp` with the arguments in the +fiber products flipped. -/ +@[simps (config := .lemmasOnly)] +def Cover.pullbackCoverOverProp' : W.Cover P where + J := 𝒰.J + obj x := (pullback ((𝒰.map x).asOverProp (hX := hQ x) (hY := hX) S) + (f.asOverProp (hX := hW) (hY := hX) S)).left + map x := (pullback.snd ((𝒰.map x).asOverProp S) (f.asOverProp S)).left + f x := 𝒰.f (f.base x) + covers x := (mem_range_iff_of_surjective ((𝒰.pullbackCover' f).map (𝒰.f (f.base x))) _ + ((PreservesPullback.iso (MorphismProperty.Over.forget Q _ _ ⋙ Over.forget S) + ((𝒰.map _).asOverProp S) (f.asOverProp S)).inv) + (PreservesPullback.iso_inv_snd _ _ _) x).mp ((𝒰.pullbackCover' f).covers x) + map_prop j := by + dsimp only + rw [← Over.forget_map, MorphismProperty.Comma.toCommaMorphism_eq_hom, + ← MorphismProperty.Comma.forget_map, ← Functor.comp_map] + rw [← PreservesPullback.iso_hom_snd, P.cancel_left_of_respectsIso] + exact P.pullback_snd _ _ (𝒰.map_prop j) + +instance (j : 𝒰.J) : ((𝒰.pullbackCoverOverProp' S f hX hW hQ).obj j).Over S where + hom := (pullback ((𝒰.map j).asOverProp (hX := hQ j) (hY := hX) S) + (f.asOverProp (hX := hW) (hY := hX) S)).hom + +instance : (𝒰.pullbackCoverOverProp' S f hX hW hQ).Over S where + isOver_map j := + { comp_over := by exact (pullback.snd ((𝒰.map j).asOverProp S) (f.asOverProp S)).w } + +end + +variable [P.IsStableUnderComposition] +variable {X : Scheme.{u}} (𝒰 : X.Cover P) (𝒱 : ∀ x, (𝒰.obj x).Cover P) + [X.Over S] [𝒰.Over S] [∀ x, (𝒱 x).Over S] + +instance (j : (𝒰.bind 𝒱).J) : ((𝒰.bind 𝒱).obj j).Over S := + inferInstanceAs <| ((𝒱 j.1).obj j.2).Over S + +instance {X : Scheme.{u}} (𝒰 : X.Cover P) (𝒱 : ∀ x, (𝒰.obj x).Cover P) + [X.Over S] [𝒰.Over S] [∀ x, (𝒱 x).Over S] : (𝒰.bind 𝒱).Over S where + over := fun ⟨i, j⟩ ↦ inferInstanceAs <| ((𝒱 i).obj j).Over S + isOver_map := fun ⟨i, j⟩ ↦ { comp_over := by simp } + +end AlgebraicGeometry.Scheme diff --git a/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean b/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean index dbcc339026d05..065d25bd88b11 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean @@ -83,6 +83,18 @@ instance surjective_isLocalAtTarget : IsLocalAtTarget @Surjective := by obtain ⟨⟨y, _⟩, hy⟩ := hf i ⟨x, hxi⟩ exact ⟨y, congr(($hy).1)⟩ +@[simp] +lemma range_eq_univ [Surjective f] : Set.range f.base = Set.univ := by + simpa [Set.range_eq_univ] using f.surjective + +lemma range_eq_range_of_surjective {S : Scheme.{u}} (f : X ⟶ S) (g : Y ⟶ S) (e : X ⟶ Y) + [Surjective e] (hge : e ≫ g = f) : Set.range f.base = Set.range g.base := by + rw [← hge] + simp [Set.range_comp] + +lemma mem_range_iff_of_surjective {S : Scheme.{u}} (f : X ⟶ S) (g : Y ⟶ S) (e : X ⟶ Y) + [Surjective e] (hge : e ≫ g = f) (s : S) : s ∈ Set.range f.base ↔ s ∈ Set.range g.base := by + rw [range_eq_range_of_surjective f g e hge] end Surjective section Injective diff --git a/Mathlib/AlgebraicGeometry/Over.lean b/Mathlib/AlgebraicGeometry/Over.lean index b0f3bcaef6d0e..15349f1b13a3e 100644 --- a/Mathlib/AlgebraicGeometry/Over.lean +++ b/Mathlib/AlgebraicGeometry/Over.lean @@ -47,4 +47,11 @@ lemma Hom.isOver_iff [X.Over S] [Y.Over S] {f : X ⟶ Y} : f.IsOver S ↔ f ≫ /-! Also note the existence of `CategoryTheory.IsOverTower X Y S`. -/ +/-- Given `X.Over S`, this is the bundled object of `Over S`. -/ +abbrev asOver (X S : Scheme.{u}) [X.Over S] := OverClass.asOver X S + +/-- Given a morphism `X ⟶ Y` with `f.IsOver S`, this is the bundled morphism in `Over S`. -/ +abbrev Hom.asOver (f : X.Hom Y) (S : Scheme.{u}) [X.Over S] [Y.Over S] [f.IsOver S] := + OverClass.asOverHom S f + end AlgebraicGeometry.Scheme diff --git a/Mathlib/CategoryTheory/Comma/OverClass.lean b/Mathlib/CategoryTheory/Comma/OverClass.lean index c9ad20a1fd78e..a6b1fc35caeb2 100644 --- a/Mathlib/CategoryTheory/Comma/OverClass.lean +++ b/Mathlib/CategoryTheory/Comma/OverClass.lean @@ -120,7 +120,22 @@ instance [OverClass X S] [IsOverTower X S S'] [IsOverTower Y S S'] [HomIsOver f S] : HomIsOver f S' := homIsOver_of_isOverTower f S S' +variable (X) in /-- Bundle `X` with an `OverClass X S` instance into `Over S`. -/ +@[simps! hom left] def OverClass.asOver [OverClass X S] : Over S := Over.mk (X ↘ S) +/-- Bundle a morphism `f : X ⟶ Y` with `HomIsOver f S` into a morphism in `Over S`. -/ +@[simps! left] +def OverClass.asOverHom [OverClass X S] [OverClass Y S] (f : X ⟶ Y) [HomIsOver f S] : + OverClass.asOver X S ⟶ OverClass.asOver Y S := + Over.homMk f (comp_over f S) + +@[simps] +instance OverClass.fromOver {S : C} (X : Over S) : OverClass X.left S where + hom := X.hom + +instance {S : C} {X Y : Over S} (f : X ⟶ Y) : HomIsOver f.left S where + comp_over := Over.w f + end CategoryTheory diff --git a/Mathlib/CategoryTheory/Limits/MorphismProperty.lean b/Mathlib/CategoryTheory/Limits/MorphismProperty.lean index 8f764982cf362..b4bb842c72dd2 100644 --- a/Mathlib/CategoryTheory/Limits/MorphismProperty.lean +++ b/Mathlib/CategoryTheory/Limits/MorphismProperty.lean @@ -134,7 +134,7 @@ instance [P.ContainsIdentities] : HasTerminal (P.Over ⊤ X) := /-- If `P` is stable under composition, base change and satisfies post-cancellation, `Over.forget P ⊤ X` creates pullbacks. -/ -noncomputable def createsLimitsOfShape_walkingCospan [HasPullbacks T] +noncomputable instance createsLimitsOfShape_walkingCospan [HasPullbacks T] [P.IsStableUnderComposition] [P.IsStableUnderBaseChange] [P.HasOfPostcompProperty P] : CreatesLimitsOfShape WalkingCospan (Over.forget P ⊤ X) := haveI : HasLimitsOfShape WalkingCospan (Comma (𝟭 T) (Functor.fromPUnit X)) := diff --git a/Mathlib/CategoryTheory/MorphismProperty/Comma.lean b/Mathlib/CategoryTheory/MorphismProperty/Comma.lean index 0662c8046f120..6b093565af86f 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/Comma.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/Comma.lean @@ -117,6 +117,8 @@ instance : Category (P.Comma L R Q W) where id X := X.id comp f g := f.comp g +lemma toCommaMorphism_eq_hom {X Y : P.Comma L R Q W} (f : X ⟶ Y) : f.toCommaMorphism = f.hom := rfl + /-- Alternative `ext` lemma for `Comma.Hom`. -/ @[ext] lemma Hom.ext' {X Y : P.Comma L R Q W} {f g : X ⟶ Y} (h : f.hom = g.hom) : @@ -220,6 +222,9 @@ protected abbrev Over : Type _ := protected abbrev Over.forget : P.Over Q X ⥤ Over X := Comma.forget (Functor.id T) (Functor.fromPUnit.{0} X) P Q ⊤ +instance : (Over.forget P ⊤ X).Faithful := inferInstanceAs <| (Comma.forget _ _ _ _ _).Faithful +instance : (Over.forget P ⊤ X).Full := inferInstanceAs <| (Comma.forget _ _ _ _ _).Full + variable {P Q X} /-- Construct a morphism in `P.Over Q X` from a morphism in `Over.X`. -/ @@ -261,6 +266,9 @@ protected abbrev Under : Type _ := protected abbrev Under.forget : P.Under Q X ⥤ Under X := Comma.forget (Functor.fromPUnit.{0} X) (Functor.id T) P ⊤ Q +instance : (Under.forget P ⊤ X).Faithful := inferInstanceAs <| (Comma.forget _ _ _ _ _).Faithful +instance : (Under.forget P ⊤ X).Full := inferInstanceAs <| (Comma.forget _ _ _ _ _).Full + variable {P Q X} /-- Construct a morphism in `P.Under Q X` from a morphism in `Under.X`. -/ From 551481f4b8470d7f67ced27d202a86942a642a49 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Mon, 25 Nov 2024 16:30:32 +0000 Subject: [PATCH 296/829] feat: ring with discrete PrimeSpectrum (#18891) Prove two characterizations of commutative semirings R whose prime spectra have discrete topology: + R has finitely many prime ideals, and every prime ideal is maximal; + R has finitely many maximal ideals, the intersection of which (= Jacobson radical) is contained in (and therefore equal to) the nilradical. From the first characterization it is easy to show the spectrum of an Artinian ring has discrete topology, but Artinian doesn't currently import PrimeSpectrum. Also show that if R is a commutative ring and Spec R has discrete topology, then the canonical RingHom from R to the product of localizations at maximal ideals (which is always injective) is an isomorphism. This is not a characterization, because it is satisfied by every local ring. (But it's a characterization if we replace "maximal" by "prime".) For this purpose we prove that, if a basic open D(f) in Spec R is a singleton {p}, then localization away from f is the same as localization at p. A key fact used is that if an ideal in a commutative semiring is disjoint from a submonoid, then there exists a maximal such ideal containing the ideal, and any maximal such ideal is prime. Also upgrade `exists_idempotent_basicOpen_eq_of_isClopen` to `existsUnique`. --- .../Algebra/Group/Submonoid/Membership.lean | 16 ++ Mathlib/Algebra/Ring/Idempotents.lean | 11 + .../PrimeSpectrum/Basic.lean | 196 ++++++++++++++++-- Mathlib/RingTheory/Artinian.lean | 7 +- Mathlib/RingTheory/Ideal/Maximal.lean | 20 ++ Mathlib/RingTheory/Ideal/Span.lean | 19 ++ Mathlib/RingTheory/Idempotents.lean | 2 + Mathlib/RingTheory/Localization/AtPrime.lean | 20 ++ .../RingTheory/Localization/Away/Basic.lean | 57 ++--- .../RingTheory/Localization/Away/Lemmas.lean | 10 +- Mathlib/RingTheory/Localization/Basic.lean | 12 +- Mathlib/RingTheory/Localization/Defs.lean | 5 + Mathlib/RingTheory/Nilpotent/Basic.lean | 8 + Mathlib/RingTheory/PrimeSpectrum.lean | 8 + Mathlib/Topology/Order.lean | 5 + 15 files changed, 337 insertions(+), 59 deletions(-) diff --git a/Mathlib/Algebra/Group/Submonoid/Membership.lean b/Mathlib/Algebra/Group/Submonoid/Membership.lean index 2f07619bf80ce..b48b123c3e191 100644 --- a/Mathlib/Algebra/Group/Submonoid/Membership.lean +++ b/Mathlib/Algebra/Group/Submonoid/Membership.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.FreeMonoid.Basic import Mathlib.Algebra.Group.Submonoid.MulOpposite import Mathlib.Algebra.Group.Submonoid.Operations import Mathlib.Algebra.GroupWithZero.Divisibility +import Mathlib.Algebra.Ring.Idempotents import Mathlib.Algebra.Ring.Int.Defs import Mathlib.Data.Finset.NoncommProd import Mathlib.Data.Nat.Cast.Basic @@ -420,6 +421,21 @@ lemma powers_le {n : M} {P : Submonoid M} : powers n ≤ P ↔ n ∈ P := by sim lemma powers_one : powers (1 : M) = ⊥ := bot_unique <| powers_le.2 <| one_mem _ +theorem _root_.IsIdempotentElem.coe_powers {a : M} (ha : IsIdempotentElem a) : + (Submonoid.powers a : Set M) = {1, a} := + let S : Submonoid M := + { carrier := {1, a}, + mul_mem' := by + rintro _ _ (rfl|rfl) (rfl|rfl) + · rw [one_mul]; exact .inl rfl + · rw [one_mul]; exact .inr rfl + · rw [mul_one]; exact .inr rfl + · rw [ha]; exact .inr rfl + one_mem' := .inl rfl } + suffices Submonoid.powers a = S from congr_arg _ this + le_antisymm (Submonoid.powers_le.mpr <| .inr rfl) + (by rintro _ (rfl|rfl); exacts [one_mem _, Submonoid.mem_powers _]) + /-- The submonoid generated by an element is a group if that element has finite order. -/ abbrev groupPowers {x : M} {n : ℕ} (hpos : 0 < n) (hx : x ^ n = 1) : Group (powers x) where inv x := x ^ (n - 1) diff --git a/Mathlib/Algebra/Ring/Idempotents.lean b/Mathlib/Algebra/Ring/Idempotents.lean index 80a2c2e0be568..8b7ce4e91aadd 100644 --- a/Mathlib/Algebra/Ring/Idempotents.lean +++ b/Mathlib/Algebra/Ring/Idempotents.lean @@ -67,6 +67,17 @@ theorem one_sub {p : R} (h : IsIdempotentElem p) : IsIdempotentElem (1 - p) := b theorem one_sub_iff {p : R} : IsIdempotentElem (1 - p) ↔ IsIdempotentElem p := ⟨fun h => sub_sub_cancel 1 p ▸ h.one_sub, IsIdempotentElem.one_sub⟩ +theorem add_sub_mul_of_commute {R} [Ring R] {p q : R} (h : Commute p q) + (hp : IsIdempotentElem p) (hq : IsIdempotentElem q) : + IsIdempotentElem (p + q - p * q) := by + convert (hp.one_sub.mul_of_commute ?_ hq.one_sub).one_sub using 1 + · simp_rw [sub_mul, mul_sub, one_mul, mul_one, sub_sub, sub_sub_cancel, add_sub, add_comm] + · simp_rw [commute_iff_eq, sub_mul, mul_sub, one_mul, mul_one, sub_sub, add_sub, add_comm, h.eq] + +theorem add_sub_mul {R} [CommRing R] {p q : R} (hp : IsIdempotentElem p) (hq : IsIdempotentElem q) : + IsIdempotentElem (p + q - p * q) := + add_sub_mul_of_commute (mul_comm p q) hp hq + theorem pow {p : N} (n : ℕ) (h : IsIdempotentElem p) : IsIdempotentElem (p ^ n) := Nat.recOn n ((pow_zero p).symm ▸ one) fun n _ => show p ^ n.succ * p ^ n.succ = p ^ n.succ by diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index 8121592d7b482..6a4899d2598b7 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -9,7 +9,7 @@ import Mathlib.RingTheory.Ideal.Over import Mathlib.RingTheory.KrullDimension.Basic import Mathlib.RingTheory.LocalRing.ResidueField.Defs import Mathlib.RingTheory.LocalRing.RingHom.Basic -import Mathlib.RingTheory.Localization.Away.Basic +import Mathlib.RingTheory.Localization.Away.Lemmas import Mathlib.Tactic.StacksAttribute import Mathlib.Topology.KrullDimension import Mathlib.Topology.Sober @@ -212,6 +212,31 @@ instance compactSpace : CompactSpace (PrimeSpectrum R) := by simp_rw [hI, ← zeroLocus_iSup, zeroLocus_empty_iff_eq_top, ← top_le_iff] at S_empty ⊢ exact Ideal.isCompactElement_top.exists_finset_of_le_iSup _ _ S_empty +/-- The prime spectrum of a semiring has discrete Zariski topology iff it is finite and +all primes are maximal. -/ +theorem discreteTopology_iff_finite_and_isPrime_imp_isMaximal : DiscreteTopology (PrimeSpectrum R) ↔ + Finite (PrimeSpectrum R) ∧ ∀ I : Ideal R, I.IsPrime → I.IsMaximal := + ⟨fun _ ↦ ⟨finite_of_compact_of_discrete, fun I hI ↦ (isClosed_singleton_iff_isMaximal ⟨I, hI⟩).mp + <| discreteTopology_iff_forall_isClosed.mp ‹_› _⟩, fun ⟨_, h⟩ ↦ .of_finite_of_isClosed_singleton + fun p ↦ (isClosed_singleton_iff_isMaximal p).mpr <| h _ p.2⟩ + +/-- The prime spectrum of a semiring has discrete Zariski topology iff there are only +finitely many maximal ideals and their intersection is contained in the nilradical. -/ +theorem discreteTopology_iff_finite_isMaximal_and_sInf_le_nilradical : + letI s := {I : Ideal R | I.IsMaximal} + DiscreteTopology (PrimeSpectrum R) ↔ Finite s ∧ sInf s ≤ nilradical R := + discreteTopology_iff_finite_and_isPrime_imp_isMaximal.trans <| by + rw [(equivSubtype R).finite_iff, ← Set.coe_setOf, Set.finite_coe_iff, Set.finite_coe_iff] + refine ⟨fun h ↦ ⟨h.1.subset fun _ h ↦ h.isPrime, nilradical_eq_sInf R ▸ sInf_le_sInf h.2⟩, + fun ⟨fin, le⟩ ↦ ?_⟩ + have hpm (I : Ideal R) (hI : I.IsPrime): I.IsMaximal := by + replace le := le.trans (nilradical_le_prime I) + rw [← fin.coe_toFinset, ← Finset.inf_id_eq_sInf, hI.inf_le'] at le + have ⟨M, hM, hMI⟩ := le + rw [fin.mem_toFinset] at hM + rwa [← hM.eq_of_le hI.1 hMI] + exact ⟨fin.subset hpm, hpm⟩ + section Comap variable {S' : Type*} [CommSemiring S'] @@ -501,6 +526,22 @@ lemma iSup_basicOpen_eq_top_iff' {s : Set R} : conv_rhs => rw [← Subtype.range_val (s := s), ← iSup_basicOpen_eq_top_iff] simp +theorem isLocalization_away_iff_atPrime_of_basicOpen_eq_singleton [Algebra R S] + {f : R} {p : PrimeSpectrum R} (h : (basicOpen f).1 = {p}) : + IsLocalization.Away f S ↔ IsLocalization.AtPrime S p.1 := + have : IsLocalization.AtPrime (Localization.Away f) p.1 := by + refine .of_le_of_exists_dvd (Submonoid.powers f) _ + (Submonoid.powers_le.mpr <| by apply h ▸ Set.mem_singleton p) fun r hr ↦ ?_ + contrapose! hr + simp_rw [← Ideal.mem_span_singleton] at hr + have ⟨q, prime, le, disj⟩ := Ideal.exists_le_prime_disjoint (Ideal.span {r}) + (Submonoid.powers f) (Set.disjoint_right.mpr hr) + have : ⟨q, prime⟩ ∈ (basicOpen f).1 := Set.disjoint_right.mp disj (Submonoid.mem_powers f) + rw [h, Set.mem_singleton_iff] at this + rw [← this] + exact not_not.mpr (q.span_singleton_le_iff_mem.mp le) + IsLocalization.isLocalization_iff_of_isLocalization _ _ (Localization.Away f) + end BasicOpen section Order @@ -669,14 +710,30 @@ lemma isCompact_isOpen_iff {s : Set (PrimeSpectrum R)} : fun ⟨s, e⟩ ↦ ⟨s, s.finite_toSet, by simpa using e.symm⟩⟩ lemma isCompact_isOpen_iff_ideal {s : Set (PrimeSpectrum R)} : - IsCompact s ∧ IsOpen s ↔ - ∃ I : Ideal R, I.FG ∧ (PrimeSpectrum.zeroLocus (I : Set R))ᶜ = s := by + IsCompact s ∧ IsOpen s ↔ ∃ I : Ideal R, I.FG ∧ (zeroLocus I)ᶜ = s := by rw [isCompact_isOpen_iff] exact ⟨fun ⟨s, e⟩ ↦ ⟨.span s, ⟨s, rfl⟩, by simpa using e⟩, fun ⟨I, ⟨s, hs⟩, e⟩ ↦ ⟨s, by simpa [hs.symm] using e⟩⟩ -lemma exists_idempotent_basicOpen_eq_of_is_clopen {s : Set (PrimeSpectrum R)} - (hs : IsClopen s) : ∃ e : R, IsIdempotentElem e ∧ s = basicOpen e := by +lemma basicOpen_injOn_isIdempotentElem : + {e : R | IsIdempotentElem e}.InjOn basicOpen := fun x hx y hy eq ↦ by + by_contra! ne + wlog ne' : x * y ≠ x generalizing x y + · apply this y hy x hx eq.symm ne.symm + rwa [mul_comm, of_not_not ne'] + have : x ∉ Ideal.span {y} := fun mem ↦ ne' <| by + obtain ⟨r, rfl⟩ := Ideal.mem_span_singleton'.mp mem + rw [mul_assoc, hy] + have ⟨p, prime, le, nmem⟩ := Ideal.exists_le_prime_nmem_of_isIdempotentElem _ x hx this + exact ne_of_mem_of_not_mem' (a := ⟨p, prime⟩) nmem + (not_not.mpr <| p.span_singleton_le_iff_mem.mp le) eq + +@[stacks 00EE] +lemma existsUnique_idempotent_basicOpen_eq_of_isClopen {s : Set (PrimeSpectrum R)} + (hs : IsClopen s) : ∃! e : R, IsIdempotentElem e ∧ s = basicOpen e := by + refine exists_unique_of_exists_of_unique ?_ ?_; swap + · rintro x y ⟨hx, rfl⟩ ⟨hy, eq⟩ + exact basicOpen_injOn_isIdempotentElem hx hy (SetLike.ext' eq) cases subsingleton_or_nontrivial R · exact ⟨0, Subsingleton.elim _ _, Subsingleton.elim _ _⟩ obtain ⟨I, hI, hI'⟩ := isCompact_isOpen_iff_ideal.mp ⟨hs.1.isCompact, hs.2⟩ @@ -712,6 +769,13 @@ lemma exists_idempotent_basicOpen_eq_of_is_clopen {s : Set (PrimeSpectrum R)} exact PrimeSpectrum.zeroLocus_anti_mono (Set.singleton_subset_iff.mpr <| Ideal.pow_le_self hnz hx) +lemma exists_idempotent_basicOpen_eq_of_isClopen {s : Set (PrimeSpectrum R)} + (hs : IsClopen s) : ∃ e : R, IsIdempotentElem e ∧ s = basicOpen e := + (existsUnique_idempotent_basicOpen_eq_of_isClopen hs).exists + +@[deprecated (since := "2024-11-11")] +alias exists_idempotent_basicOpen_eq_of_is_clopen := exists_idempotent_basicOpen_eq_of_isClopen + section IsIntegral open Polynomial @@ -939,7 +1003,10 @@ section Idempotent variable {R} [CommRing R] -lemma PrimeSpectrum.basicOpen_eq_zeroLocus_of_isIdempotentElem +namespace PrimeSpectrum + +@[stacks 00EC] +lemma basicOpen_eq_zeroLocus_of_isIdempotentElem (e : R) (he : IsIdempotentElem e) : basicOpen e = zeroLocus {1 - e} := by ext p @@ -951,24 +1018,123 @@ lemma PrimeSpectrum.basicOpen_eq_zeroLocus_of_isIdempotentElem rw [Ideal.eq_top_iff_one, ← sub_add_cancel 1 e] exact add_mem h₁ h₂ -lemma PrimeSpectrum.zeroLocus_eq_basicOpen_of_isIdempotentElem +@[stacks 00EC] +lemma zeroLocus_eq_basicOpen_of_isIdempotentElem (e : R) (he : IsIdempotentElem e) : zeroLocus {e} = basicOpen (1 - e) := by rw [basicOpen_eq_zeroLocus_of_isIdempotentElem _ he.one_sub, sub_sub_cancel] -lemma PrimeSpectrum.isClopen_iff {s : Set (PrimeSpectrum R)} : +lemma isClopen_iff {s : Set (PrimeSpectrum R)} : IsClopen s ↔ ∃ e : R, IsIdempotentElem e ∧ s = basicOpen e := by - refine ⟨PrimeSpectrum.exists_idempotent_basicOpen_eq_of_is_clopen, ?_⟩ + refine ⟨exists_idempotent_basicOpen_eq_of_isClopen, ?_⟩ rintro ⟨e, he, rfl⟩ refine ⟨?_, (basicOpen e).2⟩ rw [PrimeSpectrum.basicOpen_eq_zeroLocus_of_isIdempotentElem e he] exact isClosed_zeroLocus _ -lemma PrimeSpectrum.isClopen_iff_zeroLocus {s : Set (PrimeSpectrum R)} : - IsClopen s ↔ ∃ e : R, IsIdempotentElem e ∧ s = zeroLocus {e} := by - rw [isClopen_iff] - refine ⟨fun ⟨e, he, h⟩ ↦ ⟨1 - e, he.one_sub, - h.trans (basicOpen_eq_zeroLocus_of_isIdempotentElem e he)⟩, fun ⟨e, he, h⟩ ↦ - ⟨1 - e, he.one_sub, h.trans (zeroLocus_eq_basicOpen_of_isIdempotentElem e he)⟩⟩ +lemma isClopen_iff_zeroLocus {s : Set (PrimeSpectrum R)} : + IsClopen s ↔ ∃ e : R, IsIdempotentElem e ∧ s = zeroLocus {e} := + isClopen_iff.trans <| ⟨fun ⟨e, he, h⟩ ↦ ⟨1 - e, he.one_sub, + h.trans (basicOpen_eq_zeroLocus_of_isIdempotentElem e he)⟩, + fun ⟨e, he, h⟩ ↦ ⟨1 - e, he.one_sub, h.trans (zeroLocus_eq_basicOpen_of_isIdempotentElem e he)⟩⟩ + +open TopologicalSpace (Clopens Opens) + +/-- Clopen subsets in the prime spectrum of a commutative ring are in 1-1 correspondence +with idempotent elements in the ring. -/ +@[stacks 00EE] +def isIdempotentElemEquivClopens : + {e : R | IsIdempotentElem e} ≃ Clopens (PrimeSpectrum R) := + .ofBijective (fun e ↦ ⟨basicOpen e.1, isClopen_iff.mpr ⟨_, e.2, rfl⟩⟩) + ⟨fun x y eq ↦ Subtype.ext (basicOpen_injOn_isIdempotentElem x.2 y.2 <| + SetLike.ext' (congr_arg (·.1) eq)), fun s ↦ + have ⟨e, he, h⟩ := exists_idempotent_basicOpen_eq_of_isClopen s.2 + ⟨⟨e, he⟩, Clopens.ext h.symm⟩⟩ + +lemma basicOpen_isIdempotentElemEquivClopens_symm (s) : + basicOpen (isIdempotentElemEquivClopens (R := R).symm s).1 = s.toOpens := + Opens.ext <| congr_arg (·.1) (isIdempotentElemEquivClopens.apply_symm_apply s) + +lemma coe_isIdempotentElemEquivClopens_apply (e) : + (isIdempotentElemEquivClopens e : Set (PrimeSpectrum R)) = basicOpen (e.1 : R) := rfl + +lemma isIdempotentElemEquivClopens_apply_toOpens (e) : + (isIdempotentElemEquivClopens e).toOpens = basicOpen (e.1 : R) := rfl + +lemma isIdempotentElemEquivClopens_mul (e₁ e₂ : {e : R | IsIdempotentElem e}) : + isIdempotentElemEquivClopens ⟨_, e₁.2.mul e₂.2⟩ = + isIdempotentElemEquivClopens e₁ ⊓ isIdempotentElemEquivClopens e₂ := + Clopens.ext <| by simp_rw [coe_isIdempotentElemEquivClopens_apply, basicOpen_mul]; rfl + +lemma isIdempotentElemEquivClopens_one_sub (e : {e : R | IsIdempotentElem e}) : + isIdempotentElemEquivClopens ⟨_, e.2.one_sub⟩ = (isIdempotentElemEquivClopens e)ᶜ := + SetLike.ext' <| by + simp_rw [Clopens.coe_compl, coe_isIdempotentElemEquivClopens_apply] + rw [basicOpen_eq_zeroLocus_compl, basicOpen_eq_zeroLocus_of_isIdempotentElem _ e.2] + +lemma isIdempotentElemEquivClopens_symm_inf (s₁ s₂) : + letI e := isIdempotentElemEquivClopens (R := R).symm + e (s₁ ⊓ s₂) = ⟨_, (e s₁).2.mul (e s₂).2⟩ := + isIdempotentElemEquivClopens.symm_apply_eq.mpr <| by + simp_rw [isIdempotentElemEquivClopens_mul, Equiv.apply_symm_apply] + +lemma isIdempotentElemEquivClopens_symm_compl (s : Clopens (PrimeSpectrum R)) : + isIdempotentElemEquivClopens.symm sᶜ = ⟨_, (isIdempotentElemEquivClopens.symm s).2.one_sub⟩ := + isIdempotentElemEquivClopens.symm_apply_eq.mpr <| by + rw [isIdempotentElemEquivClopens_one_sub, Equiv.apply_symm_apply] + +lemma isIdempotentElemEquivClopens_symm_top : + isIdempotentElemEquivClopens.symm ⊤ = ⟨(1 : R), .one⟩ := + isIdempotentElemEquivClopens.symm_apply_eq.mpr <| Clopens.ext <| by + rw [coe_isIdempotentElemEquivClopens_apply, basicOpen_one]; rfl + +lemma isIdempotentElemEquivClopens_symm_bot : + isIdempotentElemEquivClopens.symm ⊥ = ⟨(0 : R), .zero⟩ := + isIdempotentElemEquivClopens.symm_apply_eq.mpr <| Clopens.ext <| by + rw [coe_isIdempotentElemEquivClopens_apply, basicOpen_zero]; rfl + +lemma isIdempotentElemEquivClopens_symm_sup (s₁ s₂ : Clopens (PrimeSpectrum R)) : + letI e := isIdempotentElemEquivClopens (R := R).symm + e (s₁ ⊔ s₂) = ⟨_, (e s₁).2.add_sub_mul (e s₂).2⟩ := Subtype.ext <| by + rw [← compl_compl (_ ⊔ _), compl_sup, isIdempotentElemEquivClopens_symm_compl] + simp_rw [isIdempotentElemEquivClopens_symm_inf, isIdempotentElemEquivClopens_symm_compl] + ring + +end PrimeSpectrum + +variable [DiscreteTopology (PrimeSpectrum R)] +open PrimeSpectrum + +variable (R) in +lemma RingHom.toLocalizationIsMaximal_surjective_of_discreteTopology : + Function.Surjective (RingHom.toLocalizationIsMaximal R) := fun x ↦ by + let idem I := isIdempotentElemEquivClopens (R := R).symm ⟨{I}, isClopen_discrete _⟩ + let ideal I := Ideal.span {1 - (idem I).1} + let toSpec (I : {I : Ideal R | I.IsMaximal}) : PrimeSpectrum R := ⟨I.1, I.2.isPrime⟩ + have loc I : IsLocalization.AtPrime (R ⧸ ideal I) I.1 := by + rw [← isLocalization_away_iff_atPrime_of_basicOpen_eq_singleton] + exacts [IsLocalization.Away.quotient_of_isIdempotentElem (idem I).2, + congr_arg (·.1) (basicOpen_isIdempotentElemEquivClopens_symm ⟨{I}, isClopen_discrete _⟩)] + let equiv I := IsLocalization.algEquiv I.1.primeCompl (Localization.AtPrime I.1) (R ⧸ ideal I) + have := (discreteTopology_iff_finite_isMaximal_and_sInf_le_nilradical.mp ‹_›).1 + have ⟨r, hr⟩ := Ideal.pi_quotient_surjective ?_ fun I ↦ equiv (toSpec I) (x I) + · refine ⟨r, funext fun I ↦ (equiv <| toSpec I).injective ?_⟩ + rw [← hr]; exact (equiv _).commutes r + refine fun I J ne ↦ Ideal.isCoprime_iff_exists.mpr ?_ + have := ((idem <| toSpec I).2.mul (idem <| toSpec J).2).eq_zero_of_isNilpotent <| by + simp_rw [← basicOpen_eq_bot_iff, basicOpen_mul, SetLike.ext'_iff, idem, + TopologicalSpace.Opens.coe_inf, basicOpen_isIdempotentElemEquivClopens_symm] + exact Set.singleton_inter_eq_empty.mpr fun h ↦ ne (Subtype.ext <| congr_arg (·.1) h) + simp_rw [ideal, Ideal.mem_span_singleton', exists_exists_eq_and] + exact ⟨1, idem (toSpec I), by simpa [mul_sub]⟩ + +/-- If the prime spectrum of a commutative ring R has discrete Zariski topology, then R is +canonically isomorphic to the product of its localizations at the (finitely many) maximal ideals. -/ +@[stacks 00JA +"See also `PrimeSpectrum.discreteTopology_iff_finite_isMaximal_and_sInf_le_nilradical`."] +def RingHom.toLocalizationIsMaximalEquiv : R ≃+* + Π I : {I : Ideal R // I.IsMaximal}, haveI : I.1.IsMaximal := I.2; Localization.AtPrime I.1 := + .ofBijective _ ⟨RingHom.toLocalizationIsMaximal_injective R, + RingHom.toLocalizationIsMaximal_surjective_of_discreteTopology R⟩ end Idempotent diff --git a/Mathlib/RingTheory/Artinian.lean b/Mathlib/RingTheory/Artinian.lean index f61383992490d..2911d10aa8bb9 100644 --- a/Mathlib/RingTheory/Artinian.lean +++ b/Mathlib/RingTheory/Artinian.lean @@ -10,6 +10,7 @@ import Mathlib.Order.Filter.EventuallyConst import Mathlib.RingTheory.Nakayama import Mathlib.RingTheory.SimpleModule import Mathlib.Tactic.RSuffices +import Mathlib.Tactic.StacksAttribute /-! # Artinian rings and modules @@ -347,6 +348,7 @@ end CommRing Strictly speaking, this should be called `IsLeftArtinianRing` but we omit the `Left` for convenience in the commutative case. For a right Artinian ring, use `IsArtinian Rᵐᵒᵖ R`. -/ +@[stacks 00J5] abbrev IsArtinianRing (R) [Ring R] := IsArtinian R R @@ -418,6 +420,7 @@ open IsArtinian variable {R : Type*} [CommRing R] [IsArtinianRing R] +@[stacks 00J8] theorem isNilpotent_jacobson_bot : IsNilpotent (Ideal.jacobson (⊥ : Ideal R)) := by let Jac := Ideal.jacobson (⊥ : Ideal R) let f : ℕ →o (Ideal R)ᵒᵈ := ⟨fun n => Jac ^ n, fun _ _ h => Ideal.pow_le_pow_right h⟩ @@ -504,9 +507,7 @@ lemma primeSpectrum_finite : {I : Ideal R | I.IsPrime}.Finite := by rwa [← Subtype.ext <| (@isMaximal_of_isPrime _ _ _ _ q.2).eq_of_le p.2.1 hq2] variable (R) -/-- -[Stacks Lemma 00J7](https://stacks.math.columbia.edu/tag/00J7) --/ +@[stacks 00J7] lemma maximal_ideals_finite : {I : Ideal R | I.IsMaximal}.Finite := by simp_rw [← isPrime_iff_isMaximal] apply primeSpectrum_finite R diff --git a/Mathlib/RingTheory/Ideal/Maximal.lean b/Mathlib/RingTheory/Ideal/Maximal.lean index 444898cdd4f69..c3cb20cf054d0 100644 --- a/Mathlib/RingTheory/Ideal/Maximal.lean +++ b/Mathlib/RingTheory/Ideal/Maximal.lean @@ -200,6 +200,26 @@ lemma isPrime_of_maximally_disjoint (I : Ideal α) rw [← hr₁, ← hr₂] ring +theorem exists_le_prime_disjoint (S : Submonoid α) (disjoint : Disjoint (I : Set α) S) : + ∃ p : Ideal α, p.IsPrime ∧ I ≤ p ∧ Disjoint (p : Set α) S := by + have ⟨p, hIp, hp⟩ := zorn_le_nonempty₀ {p : Ideal α | Disjoint (p : Set α) S} + (fun c hc hc' x hx ↦ ?_) I disjoint + · exact ⟨p, isPrime_of_maximally_disjoint _ _ hp.1 (fun _ ↦ hp.not_prop_of_gt), hIp, hp.1⟩ + cases isEmpty_or_nonempty c + · exact ⟨I, disjoint, fun J hJ ↦ isEmptyElim (⟨J, hJ⟩ : c)⟩ + refine ⟨sSup c, Set.disjoint_left.mpr fun x hx ↦ ?_, fun _ ↦ le_sSup⟩ + have ⟨p, hp⟩ := (Submodule.mem_iSup_of_directed _ hc'.directed).mp (sSup_eq_iSup' c ▸ hx) + exact Set.disjoint_left.mp (hc p.2) hp + +theorem exists_le_prime_nmem_of_isIdempotentElem (a : α) (ha : IsIdempotentElem a) (haI : a ∉ I) : + ∃ p : Ideal α, p.IsPrime ∧ I ≤ p ∧ a ∉ p := + have : Disjoint (I : Set α) (Submonoid.powers a) := Set.disjoint_right.mpr <| by + rw [ha.coe_powers] + rintro _ (rfl|rfl) + exacts [I.ne_top_iff_one.mp (ne_of_mem_of_not_mem' Submodule.mem_top haI).symm, haI] + have ⟨p, h1, h2, h3⟩ := exists_le_prime_disjoint _ _ this + ⟨p, h1, h2, Set.disjoint_right.mp h3 (Submonoid.mem_powers a)⟩ + end Ideal end CommSemiring diff --git a/Mathlib/RingTheory/Ideal/Span.lean b/Mathlib/RingTheory/Ideal/Span.lean index e2f263723ce31..1b10889acf12f 100644 --- a/Mathlib/RingTheory/Ideal/Span.lean +++ b/Mathlib/RingTheory/Ideal/Span.lean @@ -237,3 +237,22 @@ theorem span_singleton_neg (x : α) : (span {-x} : Ideal α) = span {x} := by end Ideal end Ring + +namespace IsIdempotentElem + +variable {R} [CommRing R] {e : R} (he : IsIdempotentElem e) +include he + +theorem ker_toSpanSingleton_eq_span : + LinearMap.ker (LinearMap.toSpanSingleton R R e) = Ideal.span {1 - e} := SetLike.ext fun x ↦ by + rw [Ideal.mem_span_singleton'] + refine ⟨fun h ↦ ⟨x, by rw [mul_sub, show x * e = 0 from h, mul_one, sub_zero]⟩, fun h ↦ ?_⟩ + obtain ⟨x, rfl⟩ := h + show x * (1 - e) * e = 0 + rw [mul_assoc, sub_mul, one_mul, he, sub_self, mul_zero] + +theorem ker_toSpanSingleton_one_sub_eq_span : + LinearMap.ker (LinearMap.toSpanSingleton R R (1 - e)) = Ideal.span {e} := by + rw [ker_toSpanSingleton_eq_span he.one_sub, sub_sub_cancel] + +end IsIdempotentElem diff --git a/Mathlib/RingTheory/Idempotents.lean b/Mathlib/RingTheory/Idempotents.lean index ffc85853ec287..65b8c4836a131 100644 --- a/Mathlib/RingTheory/Idempotents.lean +++ b/Mathlib/RingTheory/Idempotents.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.GeomSum import Mathlib.Algebra.Polynomial.AlgebraMap import Mathlib.RingTheory.Ideal.Quotient.Operations import Mathlib.RingTheory.Nilpotent.Defs +import Mathlib.Tactic.StacksAttribute /-! @@ -346,6 +347,7 @@ theorem eq_of_isNilpotent_sub_of_isIdempotentElem {e₁ e₂ : R} e₁ = e₂ := eq_of_isNilpotent_sub_of_isIdempotentElem_of_commute he₁ he₂ H (.all _ _) +@[stacks 00J9] theorem existsUnique_isIdempotentElem_eq_of_ker_isNilpotent (h : ∀ x ∈ RingHom.ker f, IsNilpotent x) (e : S) (he : e ∈ f.range) (he' : IsIdempotentElem e) : ∃! e' : R, IsIdempotentElem e' ∧ f e' = e := by diff --git a/Mathlib/RingTheory/Localization/AtPrime.lean b/Mathlib/RingTheory/Localization/AtPrime.lean index 76d9a9d2cb2e8..1337c943de11e 100644 --- a/Mathlib/RingTheory/Localization/AtPrime.lean +++ b/Mathlib/RingTheory/Localization/AtPrime.lean @@ -250,3 +250,23 @@ theorem localRingHom_comp {S : Type*} [CommSemiring S] (J : Ideal S) [hJ : J.IsP simp only [Function.comp_apply, RingHom.coe_comp, localRingHom_to_map] end Localization + +namespace RingHom + +variable (R) + +/-- The canonical ring homomorphism from a commutative semiring to the product of its +localizations at all maximal ideals. It is always injective. -/ +def toLocalizationIsMaximal : R →+* + Π I : {I : Ideal R // I.IsMaximal}, haveI : I.1.IsMaximal := I.2; Localization.AtPrime I.1 := + Pi.ringHom fun _ ↦ algebraMap R _ + +theorem toLocalizationIsMaximal_injective : + Function.Injective (RingHom.toLocalizationIsMaximal R) := fun r r' eq ↦ by + rw [← one_mul r, ← one_mul r'] + by_contra ne + have ⟨I, mI, hI⟩ := (Module.eqIdeal R r r').exists_le_maximal ((Ideal.ne_top_iff_one _).mpr ne) + have ⟨s, hs⟩ := (IsLocalization.eq_iff_exists I.primeCompl _).mp (congr_fun eq ⟨I, mI⟩) + exact s.2 (hI hs) + +end RingHom diff --git a/Mathlib/RingTheory/Localization/Away/Basic.lean b/Mathlib/RingTheory/Localization/Away/Basic.lean index 01a01a331e08b..ae069a4b51072 100644 --- a/Mathlib/RingTheory/Localization/Away/Basic.lean +++ b/Mathlib/RingTheory/Localization/Away/Basic.lean @@ -297,53 +297,42 @@ end AtUnits section Prod -lemma away_of_isIdempotentElem {R S} [CommRing R] [CommRing S] [Algebra R S] +lemma away_of_isIdempotentElem_of_mul {R S} [CommSemiring R] [CommSemiring S] [Algebra R S] {e : R} (he : IsIdempotentElem e) - (H : RingHom.ker (algebraMap R S) = Ideal.span {1 - e}) + (H : ∀ x y, algebraMap R S x = algebraMap R S y ↔ e * x = e * y) (H' : Function.Surjective (algebraMap R S)) : IsLocalization.Away e S where map_units' r := by - have : algebraMap R S e = 1 := by - rw [← (algebraMap R S).map_one, eq_comm, ← sub_eq_zero, ← map_sub, ← RingHom.mem_ker, - H, Ideal.mem_span_singleton] obtain ⟨r, n, rfl⟩ := r - simp [this] + simp [show algebraMap R S e = 1 by rw [← (algebraMap R S).map_one, H, mul_one, he]] surj' z := by obtain ⟨z, rfl⟩ := H' z exact ⟨⟨z, 1⟩, by simp⟩ - exists_of_eq {x y} h := by - rw [← sub_eq_zero, ← map_sub, ← RingHom.mem_ker, H, Ideal.mem_span_singleton] at h - obtain ⟨k, hk⟩ := h - refine ⟨⟨e, Submonoid.mem_powers e⟩, ?_⟩ - rw [← sub_eq_zero, ← mul_sub, hk, ← mul_assoc, mul_sub, mul_one, he.eq, sub_self, zero_mul] + exists_of_eq {x y} h := ⟨⟨e, Submonoid.mem_powers e⟩, (H x y).mp h⟩ -instance away_fst {R S} [CommRing R] [CommRing S] : +lemma away_of_isIdempotentElem {R S} [CommRing R] [CommRing S] [Algebra R S] + {e : R} (he : IsIdempotentElem e) + (H : RingHom.ker (algebraMap R S) = Ideal.span {1 - e}) + (H' : Function.Surjective (algebraMap R S)) : + IsLocalization.Away e S := + away_of_isIdempotentElem_of_mul he (fun x y ↦ by + rw [← sub_eq_zero, ← map_sub, ← RingHom.mem_ker, H, ← he.ker_toSpanSingleton_eq_span, + LinearMap.mem_ker, LinearMap.toSpanSingleton_apply, sub_smul, sub_eq_zero] + simp_rw [mul_comm e, smul_eq_mul]) H' + +instance away_fst {R S} [CommSemiring R] [CommSemiring S] : letI := (RingHom.fst R S).toAlgebra - IsLocalization.Away (R := R × S) (1, 0) R := by + IsLocalization.Away (R := R × S) (1, 0) R := letI := (RingHom.fst R S).toAlgebra - apply away_of_isIdempotentElem - · ext <;> simp - · ext x - simp only [RingHom.algebraMap_toAlgebra, RingHom.mem_ker, RingHom.coe_fst, - Ideal.mem_span_singleton, Prod.one_eq_mk, Prod.mk_sub_mk, sub_self, sub_zero] - constructor - · intro e; use x; ext <;> simp [e] - · rintro ⟨⟨i, j⟩, rfl⟩; simp - · exact Prod.fst_surjective - -instance away_snd {R S} [CommRing R] [CommRing S] : + away_of_isIdempotentElem_of_mul (by ext <;> simp) + (fun ⟨xR, xS⟩ ⟨yR, yS⟩ ↦ show xR = yR ↔ _ by simp) Prod.fst_surjective + +instance away_snd {R S} [CommSemiring R] [CommSemiring S] : letI := (RingHom.snd R S).toAlgebra - IsLocalization.Away (R := R × S) (0, 1) S := by + IsLocalization.Away (R := R × S) (0, 1) S := letI := (RingHom.snd R S).toAlgebra - apply away_of_isIdempotentElem - · ext <;> simp - · ext x - simp only [RingHom.algebraMap_toAlgebra, RingHom.mem_ker, RingHom.coe_snd, - Ideal.mem_span_singleton, Prod.one_eq_mk, Prod.mk_sub_mk, sub_self, sub_zero] - constructor - · intro e; use x; ext <;> simp [e] - · rintro ⟨⟨i, j⟩, rfl⟩; simp - · exact Prod.snd_surjective + away_of_isIdempotentElem_of_mul (by ext <;> simp) + (fun ⟨xR, xS⟩ ⟨yR, yS⟩ ↦ show xS = yS ↔ _ by simp) Prod.snd_surjective end Prod diff --git a/Mathlib/RingTheory/Localization/Away/Lemmas.lean b/Mathlib/RingTheory/Localization/Away/Lemmas.lean index 640763c874cbe..dc5d5600d2f7f 100644 --- a/Mathlib/RingTheory/Localization/Away/Lemmas.lean +++ b/Mathlib/RingTheory/Localization/Away/Lemmas.lean @@ -15,9 +15,7 @@ This file contains lemmas on localization away from an element requiring more im variable {R : Type*} [CommRing R] -namespace IsLocalization - -namespace Away +namespace IsLocalization.Away /-- Given a set `s` in a ring `R` and for every `t : s` a set `p t` of fractions in a localization of `R` at `t`, this is the function sending a pair `(t, y)`, with @@ -61,6 +59,8 @@ lemma span_range_mulNumerator_eq_top {s : Set R} use n + m simpa [pow_add, hc] using Ideal.mul_mem_left _ _ hy -end Away +lemma quotient_of_isIdempotentElem {e : R} (he : IsIdempotentElem e) : + IsLocalization.Away e (R ⧸ Ideal.span {1 - e}) := + away_of_isIdempotentElem he Ideal.mk_ker Quotient.mk_surjective -end IsLocalization +end IsLocalization.Away diff --git a/Mathlib/RingTheory/Localization/Basic.lean b/Mathlib/RingTheory/Localization/Basic.lean index f2d6adffdcf45..9159d55e9cdbd 100644 --- a/Mathlib/RingTheory/Localization/Basic.lean +++ b/Mathlib/RingTheory/Localization/Basic.lean @@ -73,7 +73,7 @@ open Function section CommSemiring -variable {R : Type*} [CommSemiring R] {M : Submonoid R} {S : Type*} [CommSemiring S] +variable {R : Type*} [CommSemiring R] {M N : Submonoid R} {S : Type*} [CommSemiring S] variable [Algebra R S] {P : Type*} [CommSemiring P] namespace IsLocalization @@ -237,7 +237,7 @@ end IsLocalization section -variable (M) +variable (M N) theorem isLocalization_of_algEquiv [Algebra R P] [IsLocalization M S] (h : S ≃ₐ[R] P) : IsLocalization M P := by @@ -265,6 +265,14 @@ theorem isLocalization_iff_of_ringEquiv (h : S ≃+* P) : letI := (h.toRingHom.comp <| algebraMap R S).toAlgebra isLocalization_iff_of_algEquiv M { h with commutes' := fun _ => rfl } +variable (S) in +/-- If an algebra is simultaneously localizations for two submonoids, then an arbitrary algebra +is a localization of one submonoid iff it is a localization of the other. -/ +theorem isLocalization_iff_of_isLocalization [IsLocalization M S] [IsLocalization N S] + [Algebra R P] : IsLocalization M P ↔ IsLocalization N P := + ⟨fun _ ↦ isLocalization_of_algEquiv N (algEquiv M S P), + fun _ ↦ isLocalization_of_algEquiv M (algEquiv N S P)⟩ + end variable (M) diff --git a/Mathlib/RingTheory/Localization/Defs.lean b/Mathlib/RingTheory/Localization/Defs.lean index 486cd6a8eb03e..60509638366d3 100644 --- a/Mathlib/RingTheory/Localization/Defs.lean +++ b/Mathlib/RingTheory/Localization/Defs.lean @@ -140,6 +140,11 @@ theorem of_le (N : Submonoid R) (h₁ : M ≤ N) (h₂ : ∀ r ∈ N, IsUnit (al rintro ⟨c, hc⟩ exact ⟨⟨c, h₁ c.2⟩, hc⟩ +theorem of_le_of_exists_dvd (N : Submonoid R) (h₁ : M ≤ N) (h₂ : ∀ n ∈ N, ∃ m ∈ M, n ∣ m) : + IsLocalization N S := + of_le M N h₁ fun n hn ↦ have ⟨m, hm, dvd⟩ := h₂ n hn + isUnit_of_dvd_unit (map_dvd _ dvd) (map_units S ⟨m, hm⟩) + variable (S) /-- `IsLocalization.toLocalizationWithZeroMap M S` shows `S` is the monoid localization of diff --git a/Mathlib/RingTheory/Nilpotent/Basic.lean b/Mathlib/RingTheory/Nilpotent/Basic.lean index 9eddb6417010a..602de4bcb8dbd 100644 --- a/Mathlib/RingTheory/Nilpotent/Basic.lean +++ b/Mathlib/RingTheory/Nilpotent/Basic.lean @@ -91,6 +91,14 @@ lemma IsNilpotent.not_isUnit [Ring R] [Nontrivial R] {x : R} (hx : IsNilpotent x ¬ IsUnit x := mt IsUnit.not_isNilpotent (by simpa only [not_not] using hx) +lemma IsIdempotentElem.eq_zero_of_isNilpotent [MonoidWithZero R] {e : R} + (idem : IsIdempotentElem e) (nilp : IsNilpotent e) : e = 0 := by + obtain ⟨rfl | n, hn⟩ := nilp + · rw [pow_zero] at hn; rw [← one_mul e, hn, zero_mul] + · rw [← hn, idem.pow_succ_eq] + +alias IsNilpotent.eq_zero_of_isIdempotentElem := IsIdempotentElem.eq_zero_of_isNilpotent + instance [Zero R] [Pow R ℕ] [Zero S] [Pow S ℕ] [IsReduced R] [IsReduced S] : IsReduced (R × S) where eq_zero _ := fun ⟨n, hn⟩ ↦ have hn := Prod.ext_iff.1 hn Prod.ext (IsReduced.eq_zero _ ⟨n, hn.1⟩) (IsReduced.eq_zero _ ⟨n, hn.2⟩) diff --git a/Mathlib/RingTheory/PrimeSpectrum.lean b/Mathlib/RingTheory/PrimeSpectrum.lean index dd7f6aa053423..5ac44c10a7f5f 100644 --- a/Mathlib/RingTheory/PrimeSpectrum.lean +++ b/Mathlib/RingTheory/PrimeSpectrum.lean @@ -83,6 +83,14 @@ instance [Subsingleton R] : IsEmpty (PrimeSpectrum R) := variable (R S) +/-- The prime spectrum is in bijection with the set of prime ideals. -/ +@[simps] +def equivSubtype : PrimeSpectrum R ≃ {I : Ideal R // I.IsPrime} where + toFun I := ⟨I.asIdeal, I.2⟩ + invFun I := ⟨I, I.2⟩ + left_inv _ := rfl + right_inv _ := rfl + /-- The map from the direct sum of prime spectra to the prime spectrum of a direct product. -/ @[simp] def primeSpectrumProdOfSum : PrimeSpectrum R ⊕ PrimeSpectrum S → PrimeSpectrum (R × S) diff --git a/Mathlib/Topology/Order.lean b/Mathlib/Topology/Order.lean index 789c64a0b0511..168716335f5dc 100644 --- a/Mathlib/Topology/Order.lean +++ b/Mathlib/Topology/Order.lean @@ -305,6 +305,11 @@ theorem singletons_open_iff_discrete {X : Type*} [TopologicalSpace X] : (∀ a : X, IsOpen ({a} : Set X)) ↔ DiscreteTopology X := ⟨fun h => ⟨eq_bot_of_singletons_open h⟩, fun a _ => @isOpen_discrete _ _ a _⟩ +theorem DiscreteTopology.of_finite_of_isClosed_singleton [TopologicalSpace α] [Finite α] + (h : ∀ a : α, IsClosed {a}) : DiscreteTopology α := + discreteTopology_iff_forall_isClosed.mpr fun s ↦ + s.iUnion_of_singleton_coe ▸ isClosed_iUnion_of_finite fun _ ↦ h _ + theorem discreteTopology_iff_singleton_mem_nhds [TopologicalSpace α] : DiscreteTopology α ↔ ∀ x : α, {x} ∈ 𝓝 x := by simp only [← singletons_open_iff_discrete, isOpen_iff_mem_nhds, mem_singleton_iff, forall_eq] From 5c08dca69b1335bc2a146801d76c98a56bc24724 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Nov 2024 17:02:00 +0000 Subject: [PATCH 297/829] feat: `Finset.sup` in a group (#19460) Co-authored-by: Andrew Yang --- Mathlib.lean | 1 + Mathlib/Algebra/Order/Group/Finset.lean | 81 ++++++++++++++++++++ Mathlib/Algebra/Polynomial/Basic.lean | 3 +- Mathlib/Data/Finset/Fold.lean | 7 -- Mathlib/Data/Finset/Lattice/Fold.lean | 20 ++--- Mathlib/Order/Filter/AtTopBot/Monoid.lean | 1 + Mathlib/RingTheory/HahnSeries/Basic.lean | 1 + Mathlib/Topology/Algebra/ConstMulAction.lean | 1 + 8 files changed, 92 insertions(+), 23 deletions(-) create mode 100644 Mathlib/Algebra/Order/Group/Finset.lean diff --git a/Mathlib.lean b/Mathlib.lean index c6a763cc2c95e..344f1930eae20 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -655,6 +655,7 @@ import Mathlib.Algebra.Order.Group.CompleteLattice import Mathlib.Algebra.Order.Group.Cone import Mathlib.Algebra.Order.Group.Defs import Mathlib.Algebra.Order.Group.DenselyOrdered +import Mathlib.Algebra.Order.Group.Finset import Mathlib.Algebra.Order.Group.Indicator import Mathlib.Algebra.Order.Group.InjSurj import Mathlib.Algebra.Order.Group.Instances diff --git a/Mathlib/Algebra/Order/Group/Finset.lean b/Mathlib/Algebra/Order/Group/Finset.lean new file mode 100644 index 0000000000000..e08078aa0b808 --- /dev/null +++ b/Mathlib/Algebra/Order/Group/Finset.lean @@ -0,0 +1,81 @@ +/- +Copyright (c) 2024 Yaël Dillies, Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies, Andrew Yang +-/ +import Mathlib.Algebra.Order.Group.OrderIso +import Mathlib.Algebra.Order.Monoid.Canonical.Defs +import Mathlib.Algebra.Order.Monoid.Unbundled.MinMax +import Mathlib.Algebra.Order.Monoid.Unbundled.Pow +import Mathlib.Algebra.Order.Monoid.Unbundled.WithTop +import Mathlib.Data.Finset.Lattice.Fold + +/-! +# `Finset.sup` in a group +-/ + +assert_not_exists MonoidWithZero + +namespace Finset +variable {ι κ M G : Type*} + +lemma fold_max_add [LinearOrder M] [Add M] [AddRightMono M] (s : Finset ι) (a : WithBot M) + (f : ι → M) : s.fold max ⊥ (fun i ↦ ↑(f i) + a) = s.fold max ⊥ ((↑) ∘ f) + a := by + classical + induction' s using Finset.induction_on with a s _ ih <;> simp [*, max_add_add_right] + +@[to_additive nsmul_inf'] +lemma inf'_pow [LinearOrder M] [Monoid M] [MulLeftMono M] [MulRightMono M] (s : Finset ι) + (f : ι → M) (n : ℕ) (hs) : s.inf' hs f ^ n = s.inf' hs fun a ↦ f a ^ n := + map_finset_inf' (OrderHom.mk _ <| pow_left_mono n) hs _ + +@[to_additive nsmul_sup'] +lemma sup'_pow [LinearOrder M] [Monoid M] [MulLeftMono M] [MulRightMono M] (s : Finset ι) + (f : ι → M) (n : ℕ) (hs) : s.sup' hs f ^ n = s.sup' hs fun a ↦ f a ^ n := + map_finset_sup' (OrderHom.mk _ <| pow_left_mono n) hs _ + +section Group +variable [Group G] [LinearOrder G] + +@[to_additive "Also see `Finset.sup'_add'` that works for canonically ordered monoids."] +lemma sup'_mul [MulRightMono G] (s : Finset ι) (f : ι → G) (a : G) (hs) : + s.sup' hs f * a = s.sup' hs fun i ↦ f i * a := map_finset_sup' (OrderIso.mulRight a) hs f + +set_option linter.docPrime false in +@[to_additive "Also see `Finset.add_sup''` that works for canonically ordered monoids."] +lemma mul_sup' [MulLeftMono G] (s : Finset ι) (f : ι → G) (a : G) (hs) : + a * s.sup' hs f = s.sup' hs fun i ↦ a * f i := map_finset_sup' (OrderIso.mulLeft a) hs f + +end Group + +section CanonicallyLinearOrderedAddCommMonoid +variable [CanonicallyLinearOrderedAddCommMonoid M] [Sub M] [AddLeftReflectLE M] [OrderedSub M] + {s : Finset ι} {t : Finset κ} + +/-- Also see `Finset.sup'_add` that works for ordered groups. -/ +lemma sup'_add' (s : Finset ι) (f : ι → M) (a : M) (hs : s.Nonempty) : + s.sup' hs f + a = s.sup' hs fun i ↦ f i + a := by + apply le_antisymm + · apply add_le_of_le_tsub_right_of_le + · exact Finset.le_sup'_of_le _ hs.choose_spec le_add_self + · exact Finset.sup'_le _ _ fun i hi ↦ le_tsub_of_add_le_right (Finset.le_sup' (f · + a) hi) + · exact Finset.sup'_le _ _ fun i hi ↦ add_le_add_right (Finset.le_sup' _ hi) _ + +/-- Also see `Finset.add_sup'` that works for ordered groups. -/ +lemma add_sup'' (hs : s.Nonempty) (f : ι → M) (a : M) : + a + s.sup' hs f = s.sup' hs fun i ↦ a + f i := by simp_rw [add_comm a, Finset.sup'_add'] + +protected lemma sup_add (hs : s.Nonempty) (f : ι → M) (a : M) : + s.sup f + a = s.sup fun i ↦ f i + a := by + rw [← Finset.sup'_eq_sup hs, ← Finset.sup'_eq_sup hs, sup'_add'] + +protected lemma add_sup (hs : s.Nonempty) (f : ι → M) (a : M) : + a + s.sup f = s.sup fun i ↦ a + f i := by + rw [← Finset.sup'_eq_sup hs, ← Finset.sup'_eq_sup hs, add_sup''] + +lemma sup_add_sup (hs : s.Nonempty) (ht : t.Nonempty) (f : ι → M) (g : κ → M) : + s.sup f + t.sup g = (s ×ˢ t).sup fun ij ↦ f ij.1 + g ij.2 := by + simp only [Finset.sup_add hs, Finset.add_sup ht, Finset.sup_product_left] + +end CanonicallyLinearOrderedAddCommMonoid +end Finset diff --git a/Mathlib/Algebra/Polynomial/Basic.lean b/Mathlib/Algebra/Polynomial/Basic.lean index 56c5d26cab71c..72765a3a6e65a 100644 --- a/Mathlib/Algebra/Polynomial/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Basic.lean @@ -4,8 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Johannes Hölzl, Kim Morrison, Jens Wagemaker -/ import Mathlib.Algebra.GroupWithZero.Divisibility -import Mathlib.Data.Finset.Sort import Mathlib.Algebra.MonoidAlgebra.Defs +import Mathlib.Algebra.Order.Monoid.Unbundled.WithTop +import Mathlib.Data.Finset.Sort import Mathlib.Order.OmegaCompletePartialOrder /-! diff --git a/Mathlib/Data/Finset/Fold.lean b/Mathlib/Data/Finset/Fold.lean index 931d4b86e5a6b..fb9ed13966a4d 100644 --- a/Mathlib/Data/Finset/Fold.lean +++ b/Mathlib/Data/Finset/Fold.lean @@ -3,8 +3,6 @@ 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.Order.Monoid.Unbundled.MinMax -import Mathlib.Algebra.Order.Monoid.Unbundled.WithTop import Mathlib.Data.Finset.Image import Mathlib.Data.Multiset.Fold @@ -231,11 +229,6 @@ theorem fold_max_lt : s.fold max b f < c ↔ b < c ∧ ∀ x ∈ s, f x < c := b theorem lt_fold_max : c < s.fold max b f ↔ c < b ∨ ∃ x ∈ s, c < f x := fold_op_rel_iff_or lt_max_iff -theorem fold_max_add [Add β] [AddRightMono β] (n : WithBot β) - (s : Finset α) : (s.fold max ⊥ fun x : α => ↑(f x) + n) = s.fold max ⊥ ((↑) ∘ f) + n := by - classical - induction' s using Finset.induction_on with a s _ ih <;> simp [*, max_add_add_right] - end Order end Fold diff --git a/Mathlib/Data/Finset/Lattice/Fold.lean b/Mathlib/Data/Finset/Lattice/Fold.lean index 4c0a9c4e8b587..404b2ac9105b3 100644 --- a/Mathlib/Data/Finset/Lattice/Fold.lean +++ b/Mathlib/Data/Finset/Lattice/Fold.lean @@ -3,7 +3,6 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Mathlib.Algebra.Order.Monoid.Unbundled.Pow import Mathlib.Data.Finset.Fold import Mathlib.Data.Finset.Pi import Mathlib.Data.Finset.Prod @@ -819,13 +818,6 @@ theorem _root_.map_finset_sup' [SemilatticeSup β] [FunLike F α β] [SupHomClas f (s.sup' hs g) = s.sup' hs (f ∘ g) := by refine hs.cons_induction ?_ ?_ <;> intros <;> simp [*] -lemma nsmul_sup' {α β : Type*} [AddMonoid β] [LinearOrder β] - [AddLeftMono β] [AddRightMono β] - {s : Finset α} (hs : s.Nonempty) (f : α → β) (n : ℕ) : - s.sup' hs (fun a => n • f a) = n • s.sup' hs f := - let ns : SupHom β β := { toFun := (n • ·), map_sup' := fun _ _ => (nsmul_right_mono n).map_max } - (map_finset_sup' ns hs _).symm - /-- To rewrite from right to left, use `Finset.sup'_comp_eq_image`. -/ @[simp] theorem sup'_image [DecidableEq β] {s : Finset γ} {f : γ → β} (hs : (s.image f).Nonempty) @@ -857,6 +849,11 @@ lemma sup'_comp_eq_map {s : Finset γ} {f : γ ↪ β} (g : β → α) (hs : s.N theorem sup'_mono {s₁ s₂ : Finset β} (h : s₁ ⊆ s₂) (h₁ : s₁.Nonempty) : s₁.sup' h₁ f ≤ s₂.sup' (h₁.mono h) f := Finset.sup'_le h₁ _ (fun _ hb => le_sup' _ (h hb)) + +@[gcongr] +lemma sup'_mono_fun {hs : s.Nonempty} {f g : β → α} (h : ∀ b ∈ s, f b ≤ g b) : + s.sup' hs f ≤ s.sup' hs g := sup'_le _ _ fun b hb ↦ (h b hb).trans (le_sup' _ hb) + end Sup' section Inf' @@ -974,13 +971,6 @@ theorem _root_.map_finset_inf' [SemilatticeInf β] [FunLike F α β] [InfHomClas f (s.inf' hs g) = s.inf' hs (f ∘ g) := by refine hs.cons_induction ?_ ?_ <;> intros <;> simp [*] -lemma nsmul_inf' {α β : Type*} [AddMonoid β] [LinearOrder β] - [AddLeftMono β] [AddRightMono β] - {s : Finset α} (hs : s.Nonempty) (f : α → β) (n : ℕ) : - s.inf' hs (fun a => n • f a) = n • s.inf' hs f := - let ns : InfHom β β := { toFun := (n • ·), map_inf' := fun _ _ => (nsmul_right_mono n).map_min } - (map_finset_inf' ns hs _).symm - /-- To rewrite from right to left, use `Finset.inf'_comp_eq_image`. -/ @[simp] theorem inf'_image [DecidableEq β] {s : Finset γ} {f : γ → β} (hs : (s.image f).Nonempty) diff --git a/Mathlib/Order/Filter/AtTopBot/Monoid.lean b/Mathlib/Order/Filter/AtTopBot/Monoid.lean index 0cfa275d9a426..02cdf596f5854 100644 --- a/Mathlib/Order/Filter/AtTopBot/Monoid.lean +++ b/Mathlib/Order/Filter/AtTopBot/Monoid.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Algebra.Order.Monoid.OrderDual +import Mathlib.Algebra.Order.Monoid.Unbundled.Pow import Mathlib.Order.Filter.AtTopBot /-! diff --git a/Mathlib/RingTheory/HahnSeries/Basic.lean b/Mathlib/RingTheory/HahnSeries/Basic.lean index aa4ddc454c2dc..b074a86e1aa3b 100644 --- a/Mathlib/RingTheory/HahnSeries/Basic.lean +++ b/Mathlib/RingTheory/HahnSeries/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ import Mathlib.Algebra.Group.Support +import Mathlib.Algebra.Order.Monoid.Unbundled.WithTop import Mathlib.Order.WellFoundedSet /-! diff --git a/Mathlib/Topology/Algebra/ConstMulAction.lean b/Mathlib/Topology/Algebra/ConstMulAction.lean index 44b6433923f59..10007257a47c5 100644 --- a/Mathlib/Topology/Algebra/ConstMulAction.lean +++ b/Mathlib/Topology/Algebra/ConstMulAction.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Alex Kontorovich, Heather Macbeth -/ import Mathlib.Algebra.Module.ULift +import Mathlib.Algebra.Order.Group.Synonym import Mathlib.Data.Set.Pointwise.SMul import Mathlib.GroupTheory.GroupAction.Defs import Mathlib.Topology.Algebra.Constructions From 25781872f1cce8492209ebd11659f3db18e9c18c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Nov 2024 17:02:02 +0000 Subject: [PATCH 298/829] feat: `Finset.sup` in a group with zero (#19461) Co-authored-by: Andrew Yang --- Mathlib.lean | 2 +- .../Algebra/Order/GroupWithZero/Finset.lean | 64 +++++++++++++++++++ Mathlib/Algebra/Order/Ring/Finset.lean | 34 ---------- Mathlib/Analysis/Normed/Field/Lemmas.lean | 4 +- 4 files changed, 67 insertions(+), 37 deletions(-) create mode 100644 Mathlib/Algebra/Order/GroupWithZero/Finset.lean delete mode 100644 Mathlib/Algebra/Order/Ring/Finset.lean diff --git a/Mathlib.lean b/Mathlib.lean index 344f1930eae20..7206497b94d35 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -679,6 +679,7 @@ import Mathlib.Algebra.Order.Group.Unbundled.Int import Mathlib.Algebra.Order.Group.Units import Mathlib.Algebra.Order.GroupWithZero.Action.Synonym import Mathlib.Algebra.Order.GroupWithZero.Canonical +import Mathlib.Algebra.Order.GroupWithZero.Finset import Mathlib.Algebra.Order.GroupWithZero.Submonoid import Mathlib.Algebra.Order.GroupWithZero.Synonym import Mathlib.Algebra.Order.GroupWithZero.Unbundled @@ -739,7 +740,6 @@ import Mathlib.Algebra.Order.Ring.Canonical import Mathlib.Algebra.Order.Ring.Cast import Mathlib.Algebra.Order.Ring.Cone import Mathlib.Algebra.Order.Ring.Defs -import Mathlib.Algebra.Order.Ring.Finset import Mathlib.Algebra.Order.Ring.InjSurj import Mathlib.Algebra.Order.Ring.Int import Mathlib.Algebra.Order.Ring.Nat diff --git a/Mathlib/Algebra/Order/GroupWithZero/Finset.lean b/Mathlib/Algebra/Order/GroupWithZero/Finset.lean new file mode 100644 index 0000000000000..0047881217ebf --- /dev/null +++ b/Mathlib/Algebra/Order/GroupWithZero/Finset.lean @@ -0,0 +1,64 @@ +/- +Copyright (c) 2022 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser, Yaël Dillies, Andrew Yang +-/ +import Mathlib.Algebra.Order.GroupWithZero.Canonical +import Mathlib.Algebra.Order.GroupWithZero.Unbundled.Lemmas +import Mathlib.Data.Finset.Lattice.Fold + +/-! +# `Finset.sup` in a group with zero +-/ + +namespace Finset +variable {ι M₀ G₀ : Type*} + +section MonoidWithZero +variable [MonoidWithZero M₀] {s : Finset ι} {a b : ι → M₀} + +lemma sup_mul_le_mul_sup_of_nonneg [SemilatticeSup M₀] [OrderBot M₀] [PosMulMono M₀] [MulPosMono M₀] + (ha : ∀ i ∈ s, 0 ≤ a i) (hb : ∀ i ∈ s, 0 ≤ b i) : s.sup (a * b) ≤ s.sup a * s.sup b := + Finset.sup_le fun _i hi ↦ + mul_le_mul (le_sup hi) (le_sup hi) (hb _ hi) ((ha _ hi).trans <| le_sup hi) + +lemma mul_inf_le_inf_mul_of_nonneg [SemilatticeInf M₀] [OrderTop M₀] [PosMulMono M₀] [MulPosMono M₀] + (ha : ∀ i ∈ s, 0 ≤ a i) (hb : ∀ i ∈ s, 0 ≤ b i) : s.inf a * s.inf b ≤ s.inf (a * b) := + Finset.le_inf fun i hi ↦ mul_le_mul (inf_le hi) (inf_le hi) (Finset.le_inf hb) (ha i hi) + +lemma sup'_mul_le_mul_sup'_of_nonneg [SemilatticeSup M₀] [PosMulMono M₀] [MulPosMono M₀] + (ha : ∀ i ∈ s, 0 ≤ a i) (hb : ∀ i ∈ s, 0 ≤ b i) (hs) : + s.sup' hs (a * b) ≤ s.sup' hs a * s.sup' hs b := + sup'_le _ _ fun _i hi ↦ + mul_le_mul (le_sup' _ hi) (le_sup' _ hi) (hb _ hi) ((ha _ hi).trans <| le_sup' _ hi) + +lemma inf'_mul_le_mul_inf'_of_nonneg [SemilatticeInf M₀] [PosMulMono M₀] [MulPosMono M₀] + (ha : ∀ i ∈ s, 0 ≤ a i) (hb : ∀ i ∈ s, 0 ≤ b i) (hs) : + s.inf' hs a * s.inf' hs b ≤ s.inf' hs (a * b) := + le_inf' _ _ fun _i hi ↦ mul_le_mul (inf'_le _ hi) (inf'_le _ hi) (le_inf' _ _ hb) (ha _ hi) + +end MonoidWithZero + +section GroupWithZero +variable [GroupWithZero G₀] [SemilatticeSup G₀] {s : Finset ι} {a : G₀} + +lemma sup'_mul₀ [MulPosMono G₀] [MulPosReflectLE G₀] (ha : 0 < a) (f : ι → G₀) (s : Finset ι) (hs) : + s.sup' hs f * a = s.sup' hs fun i ↦ f i * a := map_finset_sup' (OrderIso.mulRight₀ _ ha) hs f + +set_option linter.docPrime false in +lemma mul₀_sup' [PosMulMono G₀] [PosMulReflectLE G₀] (ha : 0 < a) (f : ι → G₀) (s : Finset ι) (hs) : + a * s.sup' hs f = s.sup' hs fun i ↦ a * f i := map_finset_sup' (OrderIso.mulLeft₀ _ ha) hs f + +lemma sup'_div₀ [ZeroLEOneClass G₀] [MulPosStrictMono G₀] [MulPosReflectLE G₀] [PosMulReflectLT G₀] + (ha : 0 < a) (f : ι → G₀) (s : Finset ι) (hs) : s.sup' hs f / a = s.sup' hs fun i ↦ f i / a := + map_finset_sup' (OrderIso.divRight₀ _ ha) hs f + +end GroupWithZero + +lemma sup_div₀ [LinearOrderedCommGroupWithZero G₀] [OrderBot G₀] {a : G₀} (ha : 0 < a) + (s : Finset ι) (f : ι → G₀) : s.sup f / a = s.sup fun i ↦ f i / a := by + obtain rfl | hs := s.eq_empty_or_nonempty + · simp [← show (0 : G₀) = ⊥ from bot_unique zero_le'] + rw [← Finset.sup'_eq_sup hs, ← Finset.sup'_eq_sup hs, sup'_div₀ (ha := ha)] + +end Finset diff --git a/Mathlib/Algebra/Order/Ring/Finset.lean b/Mathlib/Algebra/Order/Ring/Finset.lean deleted file mode 100644 index dd490e170eff4..0000000000000 --- a/Mathlib/Algebra/Order/Ring/Finset.lean +++ /dev/null @@ -1,34 +0,0 @@ -/- -Copyright (c) 2022 Eric Wieser. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Eric Wieser --/ -import Mathlib.Algebra.Order.Ring.Defs -import Mathlib.Data.Finset.Lattice.Fold - -/-! -# Algebraic properties of finitary supremum --/ - -namespace Finset -variable {ι R : Type*} [LinearOrderedSemiring R] {a b : ι → R} - -theorem sup_mul_le_mul_sup_of_nonneg [OrderBot R] (s : Finset ι) (ha : ∀ i ∈ s, 0 ≤ a i) - (hb : ∀ i ∈ s, 0 ≤ b i) : s.sup (a * b) ≤ s.sup a * s.sup b := - Finset.sup_le fun _i hi ↦ - mul_le_mul (le_sup hi) (le_sup hi) (hb _ hi) ((ha _ hi).trans <| le_sup hi) - -theorem mul_inf_le_inf_mul_of_nonneg [OrderTop R] (s : Finset ι) (ha : ∀ i ∈ s, 0 ≤ a i) - (hb : ∀ i ∈ s, 0 ≤ b i) : s.inf a * s.inf b ≤ s.inf (a * b) := - Finset.le_inf fun i hi ↦ mul_le_mul (inf_le hi) (inf_le hi) (Finset.le_inf hb) (ha i hi) - -theorem sup'_mul_le_mul_sup'_of_nonneg (s : Finset ι) (H : s.Nonempty) (ha : ∀ i ∈ s, 0 ≤ a i) - (hb : ∀ i ∈ s, 0 ≤ b i) : s.sup' H (a * b) ≤ s.sup' H a * s.sup' H b := - (sup'_le _ _) fun _i hi ↦ - mul_le_mul (le_sup' _ hi) (le_sup' _ hi) (hb _ hi) ((ha _ hi).trans <| le_sup' _ hi) - -theorem inf'_mul_le_mul_inf'_of_nonneg (s : Finset ι) (H : s.Nonempty) (ha : ∀ i ∈ s, 0 ≤ a i) - (hb : ∀ i ∈ s, 0 ≤ b i) : s.inf' H a * s.inf' H b ≤ s.inf' H (a * b) := - (le_inf' _ _) fun _i hi ↦ mul_le_mul (inf'_le _ hi) (inf'_le _ hi) (le_inf' _ _ hb) (ha _ hi) - -end Finset diff --git a/Mathlib/Analysis/Normed/Field/Lemmas.lean b/Mathlib/Analysis/Normed/Field/Lemmas.lean index 88affc76dbe77..2943b39e7e097 100644 --- a/Mathlib/Analysis/Normed/Field/Lemmas.lean +++ b/Mathlib/Analysis/Normed/Field/Lemmas.lean @@ -6,7 +6,7 @@ Authors: Patrick Massot, Johannes Hölzl import Mathlib.Algebra.Group.AddChar import Mathlib.Algebra.Group.TypeTags.Finite -import Mathlib.Algebra.Order.Ring.Finset +import Mathlib.Algebra.Order.GroupWithZero.Finset import Mathlib.Analysis.Normed.Field.Basic import Mathlib.Analysis.Normed.Group.Bounded import Mathlib.Analysis.Normed.Group.Rat @@ -60,7 +60,7 @@ instance Pi.nonUnitalSeminormedRing {π : ι → Type*} [Fintype ι] Finset.univ.sup ((fun i => ‖x i‖₊) * fun i => ‖y i‖₊) := Finset.sup_mono_fun fun _ _ => norm_mul_le _ _ _ ≤ (Finset.univ.sup fun i => ‖x i‖₊) * Finset.univ.sup fun i => ‖y i‖₊ := - Finset.sup_mul_le_mul_sup_of_nonneg _ (fun _ _ => zero_le _) fun _ _ => zero_le _ + Finset.sup_mul_le_mul_sup_of_nonneg (fun _ _ => zero_le _) fun _ _ => zero_le _ } end NonUnitalSeminormedRing From dee58f263c603ed2915cb87fd3051be9ce27ce13 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Mon, 25 Nov 2024 17:36:47 +0000 Subject: [PATCH 299/829] feat(NumberTheory/LSeries/PrimesInAP): proof of Dirichlet's Theorem (#19421) Co-authored-by: Johan Commelin --- Mathlib/NumberTheory/LSeries/PrimesInAP.lean | 257 ++++++++++++++++++- docs/100.yaml | 2 + 2 files changed, 250 insertions(+), 9 deletions(-) diff --git a/Mathlib/NumberTheory/LSeries/PrimesInAP.lean b/Mathlib/NumberTheory/LSeries/PrimesInAP.lean index 6a882afa27f44..3cd0c083dd5d6 100644 --- a/Mathlib/NumberTheory/LSeries/PrimesInAP.lean +++ b/Mathlib/NumberTheory/LSeries/PrimesInAP.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Stoll -/ import Mathlib.NumberTheory.DirichletCharacter.Orthogonality -import Mathlib.NumberTheory.LSeries.DirichletContinuation import Mathlib.NumberTheory.LSeries.Linearity +import Mathlib.NumberTheory.LSeries.Nonvanishing import Mathlib.RingTheory.RootsOfUnity.AlgebraicallyClosed /-! @@ -15,12 +15,55 @@ The goal of this file is to prove **Dirichlet's Theorem**: If `q` is a positive and `a : ZMod q` is invertible, then there are infinitely many prime numbers `p` such that `(p : ZMod q) = a`. -This will be done in the following steps: - -1. Some auxiliary lemmas on infinite products and sums -2. Results on the von Mangoldt function restricted to a residue class -3. (TODO) Results on its L-series and an auxiliary function related to it -4. (TODO) Derivation of Dirichlet's Theorem +The main steps of the proof are as follows. +1. Define `ArithmeticFunction.vonMangoldt.residueClass a` for `a : ZMod q`, which is + a function `ℕ → ℝ` taking the value zero when `(n : ℤMod q) ≠ a` and `Λ n` else + (where `Λ` is the von Mangoldt function `ArithmeticFunction.vonMangoldt`; we have + `Λ (p^k) = log p` for prime powers and `Λ n = 0` otherwise.) +2. Show that this function can be written as a linear combination of functions + of the form `χ * Λ` (pointwise product) with Dirichlet characters `χ` mod `q`. + See `ArithmeticFunction.vonMangoldt.residueClass_eq`. +3. This implies that the L-series of `ArithmeticFunction.vonMangoldt.residueClass a` + agrees (on `re s > 1`) with the corresponding linear combination of negative logarithmic + derivatives of Dirichlet L-functions. + See `ArithmeticFunction.vonMangoldt.LSeries_residueClass_eq`. +4. Define an auxiliary function `ArithmeticFunction.vonMangoldt.LfunctionResidueClassAux a` that is + this linear combination of negative logarithmic derivatives of L-functions minus + `(q.totient)⁻¹/(s-1)`, which cancels the pole at `s = 1`. + See `ArithmeticFunction.vonMangoldt.eqOn_LfunctionResidueClassAux` for the statement + that the auxiliary function agrees with the L-series of + `ArithmeticFunction.vonMangoldt.residueClass` up to the term `(q.totient)⁻¹/(s-1)`. +5. Show that the auxiliary function is continuous on `re s ≥ 1`; + see `ArithmeticFunction.vonMangoldt.continuousOn_LfunctionResidueClassAux`. + This relies heavily on the non-vanishing of Dirichlet L-functions on the *closed* + half-plane `re s ≥ 1` (`DirichletCharacter.LFunction_ne_zero_of_one_le_re`), which + in turn can only be stated since we know that the L-series of a Dirichlet character + extends to an entire function (unless the character is trivial; then there is a + simple pole at `s = 1`); see `DirichletCharacter.LFunction_eq_LSeries` + (contributed by David Loeffler). +6. Show that the sum of `Λ n / n` over any residue class, but *excluding* the primes, converges. + See `ArithmeticFunction.vonMangoldt.summable_residueClass_non_primes_div`. +7. Combining these ingredients, we can deduce that the sum of `Λ n / n` over + the *primes* in a residue class must diverge. + See `ArithmeticFunction.vonMangoldt.not_summable_residueClass_prime_div`. +8. This finally easily implies that there must be infinitely many primes in the residue class. + +## Definitions + +* `ArithmeticFunction.vonMangoldt.residueClass a` (see above). +* `ArithmeticFunction.vonMangoldt.continuousOn_LfunctionResidueClassAux` (see above). + +## Main Result + +We give two versions of **Dirichlet's Theorem**: +* `Nat.setOf_prime_and_eq_mod_infinite` states that the set of primes `p` + such that `(p : ZMod q) = a` is infinite (when `a` is invertible in `ZMod q`). +* `Nat.forall_exists_prime_gt_and_eq_mod` states that for any natural number `n` + there is a prime `p > n` such that `(p : ZMod q) = a`. + +## Tags + +prime number, arithmetic progression, residue class, Dirichlet's Theorem -/ /-! @@ -214,7 +257,7 @@ lemma summable_residueClass_non_primes_div : variable [NeZero q] {a} /-- We can express `ArithmeticFunction.vonMangoldt.residueClass` as a linear combination -of twists of the von Mangoldt function with Dirichlet charaters. -/ +of twists of the von Mangoldt function by Dirichlet charaters. -/ lemma residueClass_apply (ha : IsUnit a) (n : ℕ) : residueClass a n = (q.totient : ℂ)⁻¹ * ∑ χ : DirichletCharacter ℂ q, χ a⁻¹ * χ n * vonMangoldt n := by @@ -224,7 +267,7 @@ lemma residueClass_apply (ha : IsUnit a) (n : ℕ) : ite_mul, zero_mul, ↓reduceIte, ite_self] /-- We can express `ArithmeticFunction.vonMangoldt.residueClass` as a linear combination -of twists of the von Mangoldt function with Dirichlet charaters. -/ +of twists of the von Mangoldt function by Dirichlet charaters. -/ lemma residueClass_eq (ha : IsUnit a) : ↗(residueClass a) = (q.totient : ℂ)⁻¹ • ∑ χ : DirichletCharacter ℂ q, χ a⁻¹ • (fun n : ℕ ↦ χ n * vonMangoldt n) := by @@ -248,6 +291,202 @@ lemma LSeries_residueClass_eq (ha : IsUnit a) {s : ℂ} (hs : 1 < s.re) : simp only [Pi.smul_apply, residueClass_apply ha, smul_eq_mul, ← mul_assoc, mul_inv_cancel_of_invertible, one_mul, Finset.sum_apply, Pi.mul_apply] +variable (a) + +open Classical in +/-- The auxiliary function used, e.g., with the Wiener-Ikehara Theorem to prove +Dirichlet's Theorem. On `re s > 1`, it agrees with the L-series of the von Mangoldt +function restricted to the residue class `a : ZMod q` minus the principal part +`(q.totient)⁻¹/(s-1)` of the pole at `s = 1`; +see `ArithmeticFunction.vonMangoldt.eqOn_LfunctionResidueClassAux`. -/ +noncomputable +abbrev LfunctionResidueClassAux (s : ℂ) : ℂ := + (q.totient : ℂ)⁻¹ * (-deriv (LFunctionTrivChar₁ q) s / LFunctionTrivChar₁ q s - + ∑ χ ∈ ({1}ᶜ : Finset (DirichletCharacter ℂ q)), χ a⁻¹ * deriv (LFunction χ) s / LFunction χ s) + +/-- The auxiliary function is continuous away from the zeros of the L-functions of the Dirichlet +characters mod `q` (including at `s = 1`). -/ +lemma continuousOn_LfunctionResidueClassAux' : + ContinuousOn (LfunctionResidueClassAux a) + {s | s = 1 ∨ ∀ χ : DirichletCharacter ℂ q, LFunction χ s ≠ 0} := by + rw [show LfunctionResidueClassAux a = fun s ↦ _ from rfl] + simp only [LfunctionResidueClassAux, sub_eq_add_neg] + refine continuousOn_const.mul <| ContinuousOn.add ?_ ?_ + · refine (continuousOn_neg_logDeriv_LFunctionTrivChar₁ q).mono fun s hs ↦ ?_ + have := LFunction_ne_zero_of_one_le_re (1 : DirichletCharacter ℂ q) (s := s) + simp only [ne_eq, Set.mem_setOf_eq] at hs + tauto + · simp only [← Finset.sum_neg_distrib, mul_div_assoc, ← mul_neg, ← neg_div] + refine continuousOn_finset_sum _ fun χ hχ ↦ continuousOn_const.mul ?_ + replace hχ : χ ≠ 1 := by simpa only [ne_eq, Finset.mem_compl, Finset.mem_singleton] using hχ + refine (continuousOn_neg_logDeriv_LFunction_of_nontriv hχ).mono fun s hs ↦ ?_ + simp only [ne_eq, Set.mem_setOf_eq] at hs + rcases hs with rfl | hs + · simp only [ne_eq, Set.mem_setOf_eq, one_re, le_refl, + LFunction_ne_zero_of_one_le_re χ (.inl hχ), not_false_eq_true] + · exact hs χ + +/-- The L-series of the von Mangoldt function restricted to the prime residue class `a` mod `q` +is continuous on `re s ≥ 1` except for a simple pole at `s = 1` with residue `(q.totient)⁻¹`. +The statement as given here in terms of `ArithmeticFunction.vonMangoldt.LfunctionResidueClassAux` +is equivalent. -/ +lemma continuousOn_LfunctionResidueClassAux : + ContinuousOn (LfunctionResidueClassAux a) {s | 1 ≤ s.re} := by + refine (continuousOn_LfunctionResidueClassAux' a).mono fun s hs ↦ ?_ + rcases eq_or_ne s 1 with rfl | hs₁ + · simp only [ne_eq, Set.mem_setOf_eq, true_or] + · simp only [ne_eq, Set.mem_setOf_eq, hs₁, false_or] + exact fun χ ↦ LFunction_ne_zero_of_one_le_re χ (.inr hs₁) <| Set.mem_setOf.mp hs + +variable {a} + +open scoped LSeries.notation + +/-- The auxiliary function agrees on `re s > 1` with the L-series of the von Mangoldt function +restricted to the residue class `a : ZMod q` minus the principal part `(q.totient)⁻¹/(s-1)` +of its pole at `s = 1`. -/ +lemma eqOn_LfunctionResidueClassAux (ha : IsUnit a) : + Set.EqOn (LfunctionResidueClassAux a) + (fun s ↦ L ↗(residueClass a) s - (q.totient : ℂ)⁻¹ / (s - 1)) + {s | 1 < s.re} := by + intro s hs + replace hs := Set.mem_setOf.mp hs + simp only [LSeries_residueClass_eq ha hs, LfunctionResidueClassAux] + rw [neg_div, ← neg_add', mul_neg, ← neg_mul, div_eq_mul_one_div (q.totient : ℂ)⁻¹, + sub_eq_add_neg, ← neg_mul, ← mul_add] + congrm (_ * ?_) + -- this should be easier, but `IsUnit.inv ha` does not work here + have ha' : IsUnit a⁻¹ := isUnit_of_dvd_one ⟨a, (ZMod.inv_mul_of_unit a ha).symm⟩ + classical -- for `Fintype.sum_eq_add_sum_compl` + rw [Fintype.sum_eq_add_sum_compl 1, MulChar.one_apply ha', one_mul, add_right_comm] + simp only [mul_div_assoc] + congrm (?_ + _) + have hs₁ : s ≠ 1 := fun h ↦ ((h ▸ hs).trans_eq one_re).false + rw [deriv_LFunctionTrivChar₁_apply_of_ne_one _ hs₁, LFunctionTrivChar₁, + Function.update_noteq hs₁, LFunctionTrivChar, add_div, + mul_div_mul_left _ _ (sub_ne_zero_of_ne hs₁)] + conv_lhs => enter [2, 1]; rw [← mul_one (LFunction ..)] + rw [mul_comm _ 1, mul_div_mul_right _ _ <| LFunction_ne_zero_of_one_le_re 1 (.inr hs₁) hs.le] + +/-- The auxiliary function takes real values for real arguments `x > 1`. -/ +lemma LfunctionResidueClassAux_real (ha : IsUnit a) {x : ℝ} (hx : 1 < x) : + LfunctionResidueClassAux a x = (LfunctionResidueClassAux a x).re := by + rw [eqOn_LfunctionResidueClassAux ha hx] + simp only [sub_re, ofReal_sub] + congr 1 + · rw [LSeries, re_tsum <| LSeriesSummable_of_abscissaOfAbsConv_lt_re <| + (abscissaOfAbsConv_residueClass_le_one a).trans_lt <| by norm_cast] + push_cast + refine tsum_congr fun n ↦ ?_ + rcases eq_or_ne n 0 with rfl | hn + · simp only [term_zero, zero_re, ofReal_zero] + · simp only [term_of_ne_zero hn, ← ofReal_natCast n, ← ofReal_cpow n.cast_nonneg, ← ofReal_div, + ofReal_re] + · rw [show (q.totient : ℂ) = (q.totient : ℝ) from rfl, ← ofReal_one, ← ofReal_sub, ← ofReal_inv, + ← ofReal_div, ofReal_re] + +variable {q : ℕ} [NeZero q] {a : ZMod q} + +/-- As `x` approaches `1` from the right along the real axis, the L-series of +`ArithmeticFunction.vonMangoldt.residueClass` is bounded below by `(q.totient)⁻¹/(x-1) - C`. -/ +lemma LSeries_residueClass_lower_bound (ha : IsUnit a) : + ∃ C : ℝ, ∀ {x : ℝ} (_ : x ∈ Set.Ioc 1 2), + (q.totient : ℝ)⁻¹ / (x - 1) - C ≤ ∑' n, residueClass a n / (n : ℝ) ^ x := by + have H {x : ℝ} (hx : 1 < x) : + ∑' n, residueClass a n / (n : ℝ) ^ x = + (LfunctionResidueClassAux a x).re + (q.totient : ℝ)⁻¹ / (x - 1) := by + refine ofReal_injective ?_ + simp only [ofReal_tsum, ofReal_div, ofReal_cpow (Nat.cast_nonneg _), ofReal_natCast, + ofReal_add, ofReal_inv, ofReal_sub, ofReal_one] + simp_rw [← LfunctionResidueClassAux_real ha hx, + eqOn_LfunctionResidueClassAux ha <| Set.mem_setOf.mpr (ofReal_re x ▸ hx), sub_add_cancel, + LSeries, term] + refine tsum_congr fun n ↦ ?_ + split_ifs with hn + · simp only [hn, residueClass_apply_zero, ofReal_zero, zero_div] + · rfl + have : ContinuousOn (fun x : ℝ ↦ (LfunctionResidueClassAux a x).re) (Set.Icc 1 2) := + continuous_re.continuousOn.comp (t := Set.univ) (continuousOn_LfunctionResidueClassAux a) + (fun ⦃x⦄ a ↦ trivial) |>.comp continuous_ofReal.continuousOn fun x hx ↦ by + simpa only [Set.mem_setOf_eq, ofReal_re] using hx.1 + obtain ⟨C, hC⟩ := bddBelow_def.mp <| IsCompact.bddBelow_image isCompact_Icc this + replace hC {x : ℝ} (hx : x ∈ Set.Icc 1 2) : C ≤ (LfunctionResidueClassAux a x).re := + hC (LfunctionResidueClassAux a x).re <| + Set.mem_image_of_mem (fun x : ℝ ↦ (LfunctionResidueClassAux a x).re) hx + refine ⟨-C, fun {x} hx ↦ ?_⟩ + rw [H hx.1, add_comm, sub_neg_eq_add, add_le_add_iff_left] + exact hC <| Set.mem_Icc_of_Ioc hx + +open vonMangoldt Filter Topology in +/-- The function `n ↦ Λ n / n` restricted to primes in an invertible residue class +is not summable. This then implies that there must be infinitely many such primes. -/ +lemma not_summable_residueClass_prime_div (ha : IsUnit a) : + ¬ Summable fun n : ℕ ↦ (if n.Prime then residueClass a n else 0) / n := by + intro H + have key : Summable fun n : ℕ ↦ residueClass a n / n := by + convert (summable_residueClass_non_primes_div a).add H using 2 with n + simp only [← add_div, ite_add_ite, zero_add, add_zero, ite_self] + let C := ∑' n, residueClass a n / n + have H₁ {x : ℝ} (hx : 1 < x) : ∑' n, residueClass a n / (n : ℝ) ^ x ≤ C := by + refine tsum_le_tsum (fun n ↦ ?_) ?_ key + · rcases n.eq_zero_or_pos with rfl | hn + · simp only [Nat.cast_zero, Real.zero_rpow (zero_lt_one.trans hx).ne', div_zero, le_refl] + · refine div_le_div_of_nonneg_left (residueClass_nonneg a _) (mod_cast hn) ?_ + conv_lhs => rw [← Real.rpow_one n] + exact Real.rpow_le_rpow_of_exponent_le (by norm_cast) hx.le + · exact summable_real_of_abscissaOfAbsConv_lt <| + (abscissaOfAbsConv_residueClass_le_one a).trans_lt <| mod_cast hx + obtain ⟨C', hC'⟩ := LSeries_residueClass_lower_bound ha + have H₁ {x} (hx : x ∈ Set.Ioc 1 2) : (q.totient : ℝ)⁻¹ ≤ (C + C') * (x - 1) := + (div_le_iff₀ <| sub_pos.mpr hx.1).mp <| + sub_le_iff_le_add.mp <| (hC' hx).trans (H₁ hx.1) + have hq : 0 < (q.totient : ℝ)⁻¹ := inv_pos.mpr (mod_cast q.totient.pos_of_neZero) + rcases le_or_lt (C + C') 0 with h₀ | h₀ + · have := hq.trans_le (H₁ (Set.right_mem_Ioc.mpr one_lt_two)) + rw [show (2 : ℝ) - 1 = 1 by norm_num, mul_one] at this + exact (this.trans_le h₀).false + · obtain ⟨ξ, hξ₁, hξ₂⟩ : ∃ ξ ∈ Set.Ioc 1 2, (C + C') * (ξ - 1) < (q.totient : ℝ)⁻¹ := by + refine ⟨min (1 + (q.totient : ℝ)⁻¹ / (C + C') / 2) 2, ⟨?_, min_le_right ..⟩, ?_⟩ + · simpa only [lt_inf_iff, lt_add_iff_pos_right, Nat.ofNat_pos, div_pos_iff_of_pos_right, + Nat.one_lt_ofNat, and_true] using div_pos hq h₀ + · rw [← min_sub_sub_right, add_sub_cancel_left, ← lt_div_iff₀' h₀] + exact (min_le_left ..).trans_lt <| div_lt_self (div_pos hq h₀) one_lt_two + exact ((H₁ hξ₁).trans_lt hξ₂).false + end ArithmeticFunction.vonMangoldt end arith_prog + +/-! +### Dirichlet's Theorem +-/ + +section DirichletsTheorem + +namespace Nat + +open ArithmeticFunction vonMangoldt + +variable {q : ℕ} [NeZero q] {a : ZMod q} + +/-- **Dirichlet's Theorem** on primes in arithmetic progression: if `q` is a positive +integer and `a : ZMod q` is a unit, then there are infintely many prime numbers `p` +such that `(p : ZMod q) = a`. -/ +theorem setOf_prime_and_eq_mod_infinite (ha : IsUnit a) : + {p : ℕ | p.Prime ∧ (p : ZMod q) = a}.Infinite := by + by_contra H + rw [Set.not_infinite] at H + exact not_summable_residueClass_prime_div ha <| + summable_of_finite_support <| support_residueClass_prime_div a ▸ H + +/-- **Dirichlet's Theorem** on primes in arithmetic progression: if `q` is a positive +integer and `a : ZMod q` is a unit, then there are infintely many prime numbers `p` +such that `(p : ZMod q) = a`. -/ +theorem forall_exists_prime_gt_and_eq_mod (ha : IsUnit a) (n : ℕ) : + ∃ p > n, p.Prime ∧ (p : ZMod q) = a := by + obtain ⟨p, hp₁, hp₂⟩ := Set.infinite_iff_exists_gt.mp (setOf_prime_and_eq_mod_infinite ha) n + exact ⟨p, hp₂.gt, Set.mem_setOf.mp hp₁⟩ + +end Nat + +end DirichletsTheorem diff --git a/docs/100.yaml b/docs/100.yaml index 31f5837e872ca..550cf7682b392 100644 --- a/docs/100.yaml +++ b/docs/100.yaml @@ -182,6 +182,8 @@ title : The Central Limit Theorem 48: title : Dirichlet’s Theorem + decl : Nat.setOf_prime_and_eq_mod_infinite + author : David Loeffler, Michael Stoll 49: title : The Cayley-Hamilton Theorem decl : Matrix.aeval_self_charpoly From 87489d6ce898fbcf567404c1f01e98f63bb1732f Mon Sep 17 00:00:00 2001 From: Nick Ward <102917377+gio256@users.noreply.github.com> Date: Mon, 25 Nov 2024 17:36:48 +0000 Subject: [PATCH 300/829] chore(AlgebraicTopology): create Quasicategory directory (#19463) Moves: - AlgebraicTopology.SimplicialSet.Quasicategory -> AlgebraicTopology.Quasicategory.Basic Co-Authored-By: [Johan Commelin](https://github.com/jcommelin) --- Mathlib.lean | 2 +- .../Quasicategory.lean => Quasicategory/Basic.lean} | 7 ++++--- Mathlib/AlgebraicTopology/SimplicialSet/KanComplex.lean | 2 +- Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean | 2 +- 4 files changed, 7 insertions(+), 6 deletions(-) rename Mathlib/AlgebraicTopology/{SimplicialSet/Quasicategory.lean => Quasicategory/Basic.lean} (91%) diff --git a/Mathlib.lean b/Mathlib.lean index 7206497b94d35..5785a8dab0524 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1034,6 +1034,7 @@ import Mathlib.AlgebraicTopology.FundamentalGroupoid.PUnit import Mathlib.AlgebraicTopology.FundamentalGroupoid.Product import Mathlib.AlgebraicTopology.FundamentalGroupoid.SimplyConnected import Mathlib.AlgebraicTopology.MooreComplex +import Mathlib.AlgebraicTopology.Quasicategory.Basic import Mathlib.AlgebraicTopology.SimplexCategory import Mathlib.AlgebraicTopology.SimplicialCategory.Basic import Mathlib.AlgebraicTopology.SimplicialCategory.SimplicialObject @@ -1043,7 +1044,6 @@ import Mathlib.AlgebraicTopology.SimplicialSet.KanComplex import Mathlib.AlgebraicTopology.SimplicialSet.Monoidal import Mathlib.AlgebraicTopology.SimplicialSet.Nerve import Mathlib.AlgebraicTopology.SimplicialSet.Path -import Mathlib.AlgebraicTopology.SimplicialSet.Quasicategory import Mathlib.AlgebraicTopology.SimplicialSet.StrictSegal import Mathlib.AlgebraicTopology.SingularSet import Mathlib.AlgebraicTopology.SplitSimplicialObject diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Quasicategory.lean b/Mathlib/AlgebraicTopology/Quasicategory/Basic.lean similarity index 91% rename from Mathlib/AlgebraicTopology/SimplicialSet/Quasicategory.lean rename to Mathlib/AlgebraicTopology/Quasicategory/Basic.lean index 381e65660f632..55b27d884364c 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Quasicategory.lean +++ b/Mathlib/AlgebraicTopology/Quasicategory/Basic.lean @@ -13,13 +13,14 @@ In this file we define quasicategories, a common model of infinity categories. We show that every Kan complex is a quasicategory. -In `Mathlib/AlgebraicTopology/Nerve.lean` -we show (TODO) that the nerve of a category is a quasicategory. +In `Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean` +we show that the nerve of a category is a quasicategory. ## TODO - Generalize the definition to higher universes. - See the corresponding TODO in `Mathlib/AlgebraicTopology/KanComplex.lean`. + See the corresponding TODO in + `Mathlib/AlgebraicTopology/SimplicialSet/KanComplex.lean`. -/ diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/KanComplex.lean b/Mathlib/AlgebraicTopology/SimplicialSet/KanComplex.lean index 531bc99ebc39c..87cf0ca7bf9e5 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/KanComplex.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/KanComplex.lean @@ -10,7 +10,7 @@ import Mathlib.AlgebraicTopology.SimplicialSet.Basic # Kan complexes In this file we give the definition of Kan complexes. -In `Mathlib/AlgebraicTopology/Quasicategory.lean` +In `Mathlib/AlgebraicTopology/Quasicategory/Basic.lean` we show that every Kan complex is a quasicategory. ## TODO diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean b/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean index ca5432b7bbe6b..cef9a156cab5a 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean @@ -3,9 +3,9 @@ Copyright (c) 2024 Emily Riehl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Emily Riehl, Joël Riou, Johan Commelin, Nick Ward -/ +import Mathlib.AlgebraicTopology.Quasicategory.Basic import Mathlib.AlgebraicTopology.SimplicialSet.Nerve import Mathlib.AlgebraicTopology.SimplicialSet.Path -import Mathlib.AlgebraicTopology.SimplicialSet.Quasicategory import Mathlib.CategoryTheory.Functor.KanExtension.Adjunction import Mathlib.CategoryTheory.Functor.KanExtension.Basic From f44426d3cc139d334553648bdd614cf74ca02d79 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Nov 2024 17:36:50 +0000 Subject: [PATCH 301/829] feat(FreeGroup): some basic lemmas (#19468) From LeanCamCombi --- Mathlib/GroupTheory/FreeGroup/Basic.lean | 16 ++++++++++++++-- 1 file changed, 14 insertions(+), 2 deletions(-) diff --git a/Mathlib/GroupTheory/FreeGroup/Basic.lean b/Mathlib/GroupTheory/FreeGroup/Basic.lean index d365be175ae24..2c47200b1ac4d 100644 --- a/Mathlib/GroupTheory/FreeGroup/Basic.lean +++ b/Mathlib/GroupTheory/FreeGroup/Basic.lean @@ -48,6 +48,7 @@ free group, Newman's diamond lemma, Church-Rosser theorem -/ open Relation +open scoped List universe u v w @@ -906,6 +907,9 @@ def reduce : (L : List (α × Bool)) -> List (α × Bool) := List.casesOn ih [hd1] fun hd2 tl2 => if hd1.1 = hd2.1 ∧ hd1.2 = not hd2.2 then tl2 else hd1 :: hd2 :: tl2 +@[to_additive (attr := simp)] lemma reduce_nil : reduce ([] : List (α × Bool)) = [] := rfl +@[to_additive] lemma reduce_singleton (s : α × Bool) : reduce [s] = [s] := rfl + @[to_additive (attr := simp)] theorem reduce.cons (x) : reduce (x :: L) = @@ -1099,11 +1103,19 @@ theorem reduce_invRev {w : List (α × Bool)} : reduce (invRev w) = invRev (redu have : Red (invRev (invRev w)) (invRev (reduce (invRev w))) := reduce.red.invRev rwa [invRev_invRev] at this -@[to_additive] -theorem toWord_inv {x : FreeGroup α} : x⁻¹.toWord = invRev x.toWord := by +@[to_additive (attr := simp)] +theorem toWord_inv (x : FreeGroup α) : x⁻¹.toWord = invRev x.toWord := by rcases x with ⟨L⟩ rw [quot_mk_eq_mk, inv_mk, toWord_mk, toWord_mk, reduce_invRev] +@[to_additive] +lemma toWord_mul_sublist (x y : FreeGroup α) : (x * y).toWord <+ x.toWord ++ y.toWord := by + refine Red.sublist ?_ + have : x * y = FreeGroup.mk (x.toWord ++ y.toWord) := by + rw [← FreeGroup.mul_mk, FreeGroup.mk_toWord, FreeGroup.mk_toWord] + rw [this] + exact FreeGroup.reduce.red + /-- **Constructive Church-Rosser theorem** (compare `church_rosser`). -/ @[to_additive "**Constructive Church-Rosser theorem** (compare `church_rosser`)."] def reduce.churchRosser (H12 : Red L₁ L₂) (H13 : Red L₁ L₃) : { L₄ // Red L₂ L₄ ∧ Red L₃ L₄ } := From b98172a9b91d3ffb0c8ea2f82646639724f8a3d4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Nov 2024 17:50:07 +0000 Subject: [PATCH 302/829] feat: `Icc a b + Icc c d = Icc (a + c) (b + d)` (#19398) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From GrowthInGroups Moves: * `inv_Ioi` -> `inv_Ioi₀` --- .../Order/Group/Pointwise/Interval.lean | 114 ++++++++++++------ Mathlib/Analysis/Calculus/LHopital.lean | 12 +- Mathlib/Analysis/Convex/Topology.lean | 4 +- .../Gaussian/GaussianIntegral.lean | 2 +- .../MeasureTheory/Measure/Haar/OfBasis.lean | 2 +- .../Measure/Lebesgue/Integral.lean | 4 +- Mathlib/Topology/Algebra/Order/Field.lean | 2 +- 7 files changed, 87 insertions(+), 53 deletions(-) diff --git a/Mathlib/Algebra/Order/Group/Pointwise/Interval.lean b/Mathlib/Algebra/Order/Group/Pointwise/Interval.lean index e6846ddbc75b3..fb5b4cae02b8a 100644 --- a/Mathlib/Algebra/Order/Group/Pointwise/Interval.lean +++ b/Mathlib/Algebra/Order/Group/Pointwise/Interval.lean @@ -114,6 +114,67 @@ theorem Ici_mul_Ioi_subset' (a b : α) : Ici a * Ioi b ⊆ Ioi (a * b) := by end ContravariantLT +section LinearOrderedCommMonoid +variable [LinearOrderedCommMonoid α] [MulLeftReflectLE α] [ExistsMulOfLE α] {a b c d : α} + +-- TODO: Generalise to arbitrary actions using a `smul` version of `MulLeftMono` +@[to_additive (attr := simp)] +lemma smul_Icc (a b c : α) : a • Icc b c = Icc (a * b) (a * c) := by + ext x + constructor + · rintro ⟨y, ⟨hby, hyc⟩, rfl⟩ + exact ⟨mul_le_mul_left' hby _, mul_le_mul_left' hyc _⟩ + · rintro ⟨habx, hxac⟩ + obtain ⟨y, hy, rfl⟩ := exists_one_le_mul_of_le habx + refine ⟨b * y, ⟨le_mul_of_one_le_right' hy, ?_⟩, (mul_assoc ..).symm⟩ + rwa [mul_assoc, mul_le_mul_iff_left] at hxac + +@[to_additive] +lemma Icc_mul_Icc (hab : a ≤ b) (hcd : c ≤ d) : Icc a b * Icc c d = Icc (a * c) (b * d) := by + refine (Icc_mul_Icc_subset' _ _ _ _).antisymm fun x ⟨hacx, hxbd⟩ ↦ ?_ + obtain hxbc | hbcx := le_total x (b * c) + · obtain ⟨y, hy, rfl⟩ := exists_one_le_mul_of_le hacx + refine ⟨a * y, ⟨le_mul_of_one_le_right' hy, ?_⟩, c, left_mem_Icc.2 hcd, mul_right_comm ..⟩ + rwa [mul_right_comm, mul_le_mul_iff_right] at hxbc + · obtain ⟨y, hy, rfl⟩ := exists_one_le_mul_of_le hbcx + refine ⟨b, right_mem_Icc.2 hab, c * y, ⟨le_mul_of_one_le_right' hy, ?_⟩, (mul_assoc ..).symm⟩ + rwa [mul_assoc, mul_le_mul_iff_left] at hxbd + +end LinearOrderedCommMonoid + +section OrderedCommGroup +variable [OrderedCommGroup α] + +@[to_additive (attr := simp)] lemma inv_Ici (a : α) : (Ici a)⁻¹ = Iic a⁻¹ := ext fun _x ↦ le_inv' +@[to_additive (attr := simp)] lemma inv_Iic (a : α) : (Iic a)⁻¹ = Ici a⁻¹ := ext fun _x ↦ inv_le' +@[to_additive (attr := simp)] lemma inv_Ioi (a : α) : (Ioi a)⁻¹ = Iio a⁻¹ := ext fun _x ↦ lt_inv' +@[to_additive (attr := simp)] lemma inv_Iio (a : α) : (Iio a)⁻¹ = Ioi a⁻¹ := ext fun _x ↦ inv_lt' + +@[to_additive (attr := simp)] +lemma inv_Icc (a b : α) : (Icc a b)⁻¹ = Icc b⁻¹ a⁻¹ := by simp [← Ici_inter_Iic, inter_comm] + +@[to_additive (attr := simp)] +lemma inv_Ico (a b : α) : (Ico a b)⁻¹ = Ioc b⁻¹ a⁻¹ := by + simp [← Ici_inter_Iio, ← Ioi_inter_Iic, inter_comm] + +@[to_additive (attr := simp)] +lemma inv_Ioc (a b : α) : (Ioc a b)⁻¹ = Ico b⁻¹ a⁻¹ := by + simp [← Ioi_inter_Iic, ← Ici_inter_Iio, inter_comm] + +@[to_additive (attr := simp)] +lemma inv_Ioo (a b : α) : (Ioo a b)⁻¹ = Ioo b⁻¹ a⁻¹ := by simp [← Ioi_inter_Iio, inter_comm] + +@[deprecated (since := "2024-11-23")] alias preimage_neg_Ici := neg_Ici +@[deprecated (since := "2024-11-23")] alias preimage_neg_Iic := neg_Iic +@[deprecated (since := "2024-11-23")] alias preimage_neg_Ioi := neg_Ioi +@[deprecated (since := "2024-11-23")] alias preimage_neg_Iio := neg_Iio +@[deprecated (since := "2024-11-23")] alias preimage_neg_Icc := neg_Icc +@[deprecated (since := "2024-11-23")] alias preimage_neg_Ico := neg_Ico +@[deprecated (since := "2024-11-23")] alias preimage_neg_Ioc := neg_Ioc +@[deprecated (since := "2024-11-23")] alias preimage_neg_Ioo := neg_Ioo + +end OrderedCommGroup + section OrderedAddCommGroup variable [OrderedAddCommGroup α] (a b c : α) @@ -192,41 +253,6 @@ theorem preimage_add_const_Ioc : (fun x => x + a) ⁻¹' Ioc b c = Ioc (b - a) ( theorem preimage_add_const_Ioo : (fun x => x + a) ⁻¹' Ioo b c = Ioo (b - a) (c - a) := by simp [← Ioi_inter_Iio] -/-! -### Preimages under `x ↦ -x` --/ - - -@[simp] -theorem preimage_neg_Ici : -Ici a = Iic (-a) := - ext fun _x => le_neg - -@[simp] -theorem preimage_neg_Iic : -Iic a = Ici (-a) := - ext fun _x => neg_le - -@[simp] -theorem preimage_neg_Ioi : -Ioi a = Iio (-a) := - ext fun _x => lt_neg - -@[simp] -theorem preimage_neg_Iio : -Iio a = Ioi (-a) := - ext fun _x => neg_lt - -@[simp] -theorem preimage_neg_Icc : -Icc a b = Icc (-b) (-a) := by simp [← Ici_inter_Iic, inter_comm] - -@[simp] -theorem preimage_neg_Ico : -Ico a b = Ioc (-b) (-a) := by - simp [← Ici_inter_Iio, ← Ioi_inter_Iic, inter_comm] - -@[simp] -theorem preimage_neg_Ioc : -Ioc a b = Ico (-b) (-a) := by - simp [← Ioi_inter_Iic, ← Ici_inter_Iio, inter_comm] - -@[simp] -theorem preimage_neg_Ioo : -Ioo a b = Ioo (-b) (-a) := by simp [← Ioi_inter_Iio, inter_comm] - /-! ### Preimages under `x ↦ x - a` -/ @@ -433,6 +459,15 @@ theorem Iio_add_bij : BijOn (· + a) (Iio b) (Iio (b + a)) := end OrderedAddCommGroup +section LinearOrderedCommGroup +variable [LinearOrderedCommGroup α] + +@[to_additive (attr := simp)] +lemma inv_uIcc (a b : α) : [[a, b]]⁻¹ = [[a⁻¹, b⁻¹]] := by + simp only [uIcc, inv_Icc, inv_sup, inv_inf] + +end LinearOrderedCommGroup + section LinearOrderedAddCommGroup variable [LinearOrderedAddCommGroup α] (a b c d : α) @@ -445,10 +480,9 @@ theorem preimage_const_add_uIcc : (fun x => a + x) ⁻¹' [[b, c]] = [[b - a, c theorem preimage_add_const_uIcc : (fun x => x + a) ⁻¹' [[b, c]] = [[b - a, c - a]] := by simpa only [add_comm] using preimage_const_add_uIcc a b c --- TODO: Why is the notation `-[[a, b]]` broken? -@[simp] -theorem preimage_neg_uIcc : @Neg.neg (Set α) Set.neg [[a, b]] = [[-a, -b]] := by - simp only [← Icc_min_max, preimage_neg_Icc, min_neg_neg, max_neg_neg] +@[deprecated neg_uIcc (since := "2024-11-23")] +theorem preimage_neg_uIcc : -[[a, b]] = [[-a, -b]] := by + simp only [← Icc_min_max, neg_Icc, min_neg_neg, max_neg_neg] @[simp] theorem preimage_sub_const_uIcc : (fun x => x - a) ⁻¹' [[b, c]] = [[b + a, c + a]] := by @@ -750,7 +784,7 @@ theorem inv_Ioo_0_left {a : α} (ha : 0 < a) : (Ioo 0 a)⁻¹ = Ioi a⁻¹ := by inv_inv a ▸ (inv_lt_inv₀ ((inv_pos.2 ha).trans h) (inv_pos.2 ha)).2 h⟩⟩ -theorem inv_Ioi {a : α} (ha : 0 < a) : (Ioi a)⁻¹ = Ioo 0 a⁻¹ := by +theorem inv_Ioi₀ {a : α} (ha : 0 < a) : (Ioi a)⁻¹ = Ioo 0 a⁻¹ := by rw [inv_eq_iff_eq_inv, inv_Ioo_0_left (inv_pos.2 ha), inv_inv] theorem image_const_mul_Ioi_zero {k : Type*} [LinearOrderedField k] {x : k} (hx : 0 < x) : diff --git a/Mathlib/Analysis/Calculus/LHopital.lean b/Mathlib/Analysis/Calculus/LHopital.lean index f1e691d8342cf..66b0aac6e04cc 100644 --- a/Mathlib/Analysis/Calculus/LHopital.lean +++ b/Mathlib/Analysis/Calculus/LHopital.lean @@ -110,11 +110,11 @@ theorem lhopital_zero_left_on_Ioo (hab : a < b) (hff' : ∀ x ∈ Ioo a b, HasDe comp x (hff' (-x) hx) (hasDerivAt_neg x) have hdng : ∀ x ∈ -Ioo a b, HasDerivAt (g ∘ Neg.neg) (g' (-x) * -1) x := fun x hx => comp x (hgg' (-x) hx) (hasDerivAt_neg x) - rw [preimage_neg_Ioo] at hdnf - rw [preimage_neg_Ioo] at hdng + rw [neg_Ioo] at hdnf + rw [neg_Ioo] at hdng have := lhopital_zero_right_on_Ioo (neg_lt_neg hab) hdnf hdng (by intro x hx h - apply hg' _ (by rw [← preimage_neg_Ioo] at hx; exact hx) + apply hg' _ (by rw [← neg_Ioo] at hx; exact hx) rwa [mul_comm, ← neg_eq_neg_one_mul, neg_eq_zero] at h) (hfb.comp tendsto_neg_nhdsWithin_Ioi_neg) (hgb.comp tendsto_neg_nhdsWithin_Ioi_neg) (by @@ -176,12 +176,12 @@ theorem lhopital_zero_atBot_on_Iio (hff' : ∀ x ∈ Iio a, HasDerivAt f (f' x) comp x (hff' (-x) hx) (hasDerivAt_neg x) have hdng : ∀ x ∈ -Iio a, HasDerivAt (g ∘ Neg.neg) (g' (-x) * -1) x := fun x hx => comp x (hgg' (-x) hx) (hasDerivAt_neg x) - rw [preimage_neg_Iio] at hdnf - rw [preimage_neg_Iio] at hdng + rw [neg_Iio] at hdnf + rw [neg_Iio] at hdng have := lhopital_zero_atTop_on_Ioi hdnf hdng (by intro x hx h - apply hg' _ (by rw [← preimage_neg_Iio] at hx; exact hx) + apply hg' _ (by rw [← neg_Iio] at hx; exact hx) rwa [mul_comm, ← neg_eq_neg_one_mul, neg_eq_zero] at h) (hfbot.comp tendsto_neg_atTop_atBot) (hgbot.comp tendsto_neg_atTop_atBot) (by diff --git a/Mathlib/Analysis/Convex/Topology.lean b/Mathlib/Analysis/Convex/Topology.lean index a38bd1386b8ea..86c3bddbec48e 100644 --- a/Mathlib/Analysis/Convex/Topology.lean +++ b/Mathlib/Analysis/Convex/Topology.lean @@ -360,8 +360,8 @@ theorem Convex.closure_subset_image_homothety_interior_of_one_lt {s : Set E} (hs refine ⟨homothety x t⁻¹ y, hs.openSegment_interior_closure_subset_interior hx hy ?_, (AffineEquiv.homothetyUnitsMulHom x (Units.mk0 t hne)).apply_symm_apply y⟩ - rw [openSegment_eq_image_lineMap, ← inv_one, ← inv_Ioi (zero_lt_one' ℝ), ← image_inv, image_image, - homothety_eq_lineMap] + rw [openSegment_eq_image_lineMap, ← inv_one, ← inv_Ioi₀ (zero_lt_one' ℝ), ← image_inv, + image_image, homothety_eq_lineMap] exact mem_image_of_mem _ ht /-- If we dilate a convex set about a point in its interior by a scale `t > 1`, the interior of diff --git a/Mathlib/Analysis/SpecialFunctions/Gaussian/GaussianIntegral.lean b/Mathlib/Analysis/SpecialFunctions/Gaussian/GaussianIntegral.lean index 75fc9f00bf0db..b88d8363dabf7 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gaussian/GaussianIntegral.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gaussian/GaussianIntegral.lean @@ -109,7 +109,7 @@ theorem integrable_rpow_mul_exp_neg_mul_sq {b : ℝ} (hb : 0 < b) {s : ℝ} (hs refine ⟨?_, integrableOn_rpow_mul_exp_neg_mul_sq hb hs⟩ rw [← (Measure.measurePreserving_neg (volume : Measure ℝ)).integrableOn_comp_preimage (Homeomorph.neg ℝ).measurableEmbedding] - simp only [Function.comp_def, neg_sq, neg_preimage, preimage_neg_Iio, neg_neg, neg_zero] + simp only [Function.comp_def, neg_sq, neg_preimage, neg_Iio, neg_neg, neg_zero] apply Integrable.mono' (integrableOn_rpow_mul_exp_neg_mul_sq hb hs) · apply Measurable.aestronglyMeasurable exact (measurable_id'.neg.pow measurable_const).mul diff --git a/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean b/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean index 7f76b5cbe8048..a0db8d8e207b9 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean @@ -116,7 +116,7 @@ theorem parallelepiped_orthonormalBasis_one_dim (b : OrthonormalBasis ι ℝ ℝ · right simp_rw [H, parallelepiped, Algebra.id.smul_eq_mul, A] simp only [F, Finset.univ_unique, Fin.default_eq_zero, mul_neg, mul_one, Finset.sum_neg_distrib, - Finset.sum_singleton, ← image_comp, Function.comp, image_neg, preimage_neg_Icc, neg_zero] + Finset.sum_singleton, ← image_comp, Function.comp, image_neg, neg_Icc, neg_zero] theorem parallelepiped_eq_sum_segment (v : ι → E) : parallelepiped v = ∑ i, segment ℝ 0 (v i) := by ext diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/Integral.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/Integral.lean index e4968335130b1..e72be147a0f25 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/Integral.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/Integral.lean @@ -83,7 +83,7 @@ theorem integral_comp_neg_Iic {E : Type*} [NormedAddCommGroup E] [NormedSpace (Homeomorph.neg ℝ).isClosedEmbedding.measurableEmbedding have := MeasurableEmbedding.setIntegral_map (μ := volume) A f (Ici (-c)) rw [Measure.map_neg_eq_self (volume : Measure ℝ)] at this - simp_rw [← integral_Ici_eq_integral_Ioi, this, neg_preimage, preimage_neg_Ici, neg_neg] + simp_rw [← integral_Ici_eq_integral_Ioi, this, neg_preimage, neg_Ici, neg_neg] /- @[simp] Porting note: Linter complains it does not apply to itself. Although it does apply to itself, it does not apply when `f` is more complicated -/ @@ -102,7 +102,7 @@ theorem integral_comp_abs {f : ℝ → ℝ} : rw [← Measure.map_neg_eq_self (volume : Measure ℝ)] let m : MeasurableEmbedding fun x : ℝ => -x := (Homeomorph.neg ℝ).measurableEmbedding rw [m.integrableOn_map_iff] - simp_rw [Function.comp_def, abs_neg, neg_preimage, preimage_neg_Iic, neg_zero] + simp_rw [Function.comp_def, abs_neg, neg_preimage, neg_Iic, neg_zero] exact integrableOn_Ici_iff_integrableOn_Ioi.mpr hf calc _ = (∫ x in Iic 0, f |x|) + ∫ x in Ioi 0, f |x| := by diff --git a/Mathlib/Topology/Algebra/Order/Field.lean b/Mathlib/Topology/Algebra/Order/Field.lean index ee7ea3caa6442..29f1d4890fd66 100644 --- a/Mathlib/Topology/Algebra/Order/Field.lean +++ b/Mathlib/Topology/Algebra/Order/Field.lean @@ -114,7 +114,7 @@ theorem Filter.Tendsto.neg_mul_atBot {C : 𝕜} (hC : C < 0) (hf : Tendsto f l ( @[simp] lemma inv_atTop₀ : (atTop : Filter 𝕜)⁻¹ = 𝓝[>] 0 := (((atTop_basis_Ioi' (0 : 𝕜)).map _).comp_surjective inv_surjective).eq_of_same_basis <| - (nhdsWithin_Ioi_basis _).congr (by simp) fun a ha ↦ by simp [inv_Ioi (inv_pos.2 ha)] + (nhdsWithin_Ioi_basis _).congr (by simp) fun a ha ↦ by simp [inv_Ioi₀ (inv_pos.2 ha)] @[simp] lemma inv_nhdsWithin_Ioi_zero : (𝓝[>] (0 : 𝕜))⁻¹ = atTop := by rw [← inv_atTop₀, inv_inv] From 3646342693db92ccb97e120635ea7d068d60e738 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Nov 2024 18:11:57 +0000 Subject: [PATCH 303/829] feat(MvPolynomial): more lemmas about `finSuccEquiv` (#19201) Co-authored-by: Andrew Yang --- Mathlib/Algebra/MvPolynomial/Equiv.lean | 39 ++++++++++++++++++------- Mathlib/Data/Finsupp/Fin.lean | 5 ++++ 2 files changed, 33 insertions(+), 11 deletions(-) diff --git a/Mathlib/Algebra/MvPolynomial/Equiv.lean b/Mathlib/Algebra/MvPolynomial/Equiv.lean index 6b77addf42b8d..ebcd5678aa4a4 100644 --- a/Mathlib/Algebra/MvPolynomial/Equiv.lean +++ b/Mathlib/Algebra/MvPolynomial/Equiv.lean @@ -336,7 +336,6 @@ theorem finSuccEquiv_eq : rfl · refine Fin.cases ?_ ?_ i <;> simp [finSuccEquiv] -@[simp] theorem finSuccEquiv_apply (p : MvPolynomial (Fin (n + 1)) R) : finSuccEquiv R n p = eval₂Hom (Polynomial.C.comp (C : R →+* MvPolynomial (Fin n) R)) @@ -356,10 +355,10 @@ theorem finSuccEquiv_comp_C_eq_C {R : Type u} [CommSemiring R] (n : ℕ) : variable {n} {R} -theorem finSuccEquiv_X_zero : finSuccEquiv R n (X 0) = Polynomial.X := by simp +theorem finSuccEquiv_X_zero : finSuccEquiv R n (X 0) = Polynomial.X := by simp [finSuccEquiv_apply] theorem finSuccEquiv_X_succ {j : Fin n} : finSuccEquiv R n (X j.succ) = Polynomial.C (X j) := by - simp + simp [finSuccEquiv_apply] /-- The coefficient of `m` in the `i`-th coefficient of `finSuccEquiv R n f` equals the coefficient of `Finsupp.cons i m` in `f`. -/ @@ -414,7 +413,7 @@ theorem coeff_eval_eq_eval_coeff (s' : Fin n → R) (f : Polynomial (MvPolynomia simp only [Polynomial.coeff_map] theorem support_coeff_finSuccEquiv {f : MvPolynomial (Fin (n + 1)) R} {i : ℕ} {m : Fin n →₀ ℕ} : - m ∈ (Polynomial.coeff ((finSuccEquiv R n) f) i).support ↔ Finsupp.cons i m ∈ f.support := by + m ∈ ((finSuccEquiv R n f).coeff i).support ↔ m.cons i ∈ f.support := by apply Iff.intro · intro h simpa [← finSuccEquiv_coeff_coeff] using h @@ -441,7 +440,7 @@ lemma totalDegree_coeff_finSuccEquiv_add_le (f : MvPolynomial (Fin (n + 1)) R) ( · rw [← support_coeff_finSuccEquiv] exact hσ1 -theorem finSuccEquiv_support (f : MvPolynomial (Fin (n + 1)) R) : +theorem support_finSuccEquiv (f : MvPolynomial (Fin (n + 1)) R) : (finSuccEquiv R n f).support = Finset.image (fun m : Fin (n + 1) →₀ ℕ => m 0) f.support := by ext i rw [Polynomial.mem_support_iff, Finset.mem_image, Finsupp.ne_iff] @@ -454,9 +453,14 @@ theorem finSuccEquiv_support (f : MvPolynomial (Fin (n + 1)) R) : refine ⟨tail m, ?_⟩ rwa [← coeff, zero_apply, ← mem_support_iff, support_coeff_finSuccEquiv, cons_tail] -theorem finSuccEquiv_support' {f : MvPolynomial (Fin (n + 1)) R} {i : ℕ} : - Finset.image (Finsupp.cons i) (Polynomial.coeff ((finSuccEquiv R n) f) i).support = - f.support.filter fun m => m 0 = i := by +@[deprecated (since := "2024-11-05")] alias finSuccEquiv_support := support_finSuccEquiv + +theorem mem_support_finSuccEquiv {f : MvPolynomial (Fin (n + 1)) R} {x} : + x ∈ (finSuccEquiv R n f).support ↔ x ∈ (fun m : Fin (n + 1) →₀ _ ↦ m 0) '' f.support := by + simpa using congr(x ∈ $(support_finSuccEquiv f)) + +theorem image_support_finSuccEquiv {f : MvPolynomial (Fin (n + 1)) R} {i : ℕ} : + ((finSuccEquiv R n f).coeff i).support.image (Finsupp.cons i) = {m ∈ f.support | m 0 = i} := by ext m rw [Finset.mem_filter, Finset.mem_image, mem_support_iff] conv_lhs => @@ -472,6 +476,19 @@ theorem finSuccEquiv_support' {f : MvPolynomial (Fin (n + 1)) R} {i : ℕ} : rw [← h.2, cons_tail] simp [h.1] +@[deprecated (since := "2024-11-05")] alias finSuccEquiv_support' := image_support_finSuccEquiv + +lemma mem_image_support_coeff_finSuccEquiv {f : MvPolynomial (Fin (n + 1)) R} {i : ℕ} {x} : + x ∈ Finsupp.cons i '' ((finSuccEquiv R n f).coeff i).support ↔ + x ∈ f.support ∧ x 0 = i := by + simpa using congr(x ∈ $image_support_finSuccEquiv) + +lemma mem_support_coeff_finSuccEquiv {f : MvPolynomial (Fin (n + 1)) R} {i : ℕ} {x} : + x ∈ ((finSuccEquiv R n f).coeff i).support ↔ x.cons i ∈ f.support := by + rw [← (Finsupp.cons_right_injective i).mem_finset_image (a := x), + image_support_finSuccEquiv] + simp only [Finset.mem_filter, mem_support_iff, ne_eq, cons_zero, and_true] + -- TODO: generalize `finSuccEquiv R n` to an arbitrary ZeroHom theorem support_finSuccEquiv_nonempty {f : MvPolynomial (Fin (n + 1)) R} (h : f ≠ 0) : (finSuccEquiv R n f).support.Nonempty := by @@ -485,7 +502,7 @@ theorem degree_finSuccEquiv {f : MvPolynomial (Fin (n + 1)) R} (h : f ≠ 0) : have h₂ : WithBot.some = Nat.cast := rfl have h' : ((finSuccEquiv R n f).support.sup fun x => x) = degreeOf 0 f := by - rw [degreeOf_eq_sup, finSuccEquiv_support f, Finset.sup_image, h₀] + rw [degreeOf_eq_sup, support_finSuccEquiv, Finset.sup_image, h₀] rw [Polynomial.degree, ← h', ← h₂, Finset.coe_sup_of_nonempty (support_finSuccEquiv_nonempty h), Finset.max_eq_sup_coe, h₁] @@ -524,8 +541,8 @@ lemma finSuccEquiv_rename_finSuccEquiv (e : σ ≃ Fin n) (φ : MvPolynomial (Op (Polynomial.mapRingHom (rename e).toRingHom).comp (optionEquivLeft R σ) by exact DFunLike.congr_fun this φ apply ringHom_ext - · simp [Polynomial.algebraMap_apply, algebraMap_eq] - · rintro (i|i) <;> simp + · simp [Polynomial.algebraMap_apply, algebraMap_eq, finSuccEquiv_apply] + · rintro (i|i) <;> simp [finSuccEquiv_apply] end diff --git a/Mathlib/Data/Finsupp/Fin.lean b/Mathlib/Data/Finsupp/Fin.lean index b75198493ebe5..4e7868d955fcf 100644 --- a/Mathlib/Data/Finsupp/Fin.lean +++ b/Mathlib/Data/Finsupp/Fin.lean @@ -17,6 +17,7 @@ In this context, we prove some usual properties of `tail` and `cons`, analogous `Data.Fin.Tuple.Basic`. -/ +open Function noncomputable section @@ -86,4 +87,8 @@ lemma cons_support : (s.cons y).support ⊆ insert 0 (s.support.map (Fin.succEmb rintro i rfl simpa [Finsupp.mem_support_iff] using hi +lemma cons_right_injective {n : ℕ} {M : Type*} [Zero M] (y : M) : + Injective (Finsupp.cons y : (Fin n →₀ M) → Fin (n + 1) →₀ M) := + (equivFunOnFinite.symm.injective.comp ((Fin.cons_right_injective _).comp DFunLike.coe_injective)) + end Finsupp From 69eeb7456e3ae4dbacb7ea34191d938acb9275e7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Nov 2024 20:58:34 +0000 Subject: [PATCH 304/829] chore(Data/Nat/Defs): rename two lemmas that were not about `Ne` (#19476) From LeanCamCombi --- Mathlib/Algebra/Group/Nat/Even.lean | 2 +- Mathlib/Algebra/Ring/Parity.lean | 2 +- Mathlib/Data/Nat/Defs.lean | 7 +++++-- 3 files changed, 7 insertions(+), 4 deletions(-) diff --git a/Mathlib/Algebra/Group/Nat/Even.lean b/Mathlib/Algebra/Group/Nat/Even.lean index 86fcfd3209c0d..3377025e9fb61 100644 --- a/Mathlib/Algebra/Group/Nat/Even.lean +++ b/Mathlib/Algebra/Group/Nat/Even.lean @@ -31,7 +31,7 @@ instance : DecidablePred (IsSquare : ℕ → Prop) := fun m ↦ decidable_of_iff' (Nat.sqrt m * Nat.sqrt m = m) <| by simp_rw [← Nat.exists_mul_self m, IsSquare, eq_comm] -lemma not_even_iff : ¬ Even n ↔ n % 2 = 1 := by rw [even_iff, mod_two_ne_zero] +lemma not_even_iff : ¬ Even n ↔ n % 2 = 1 := by rw [even_iff, mod_two_not_eq_zero] @[simp] lemma two_dvd_ne_zero : ¬2 ∣ n ↔ n % 2 = 1 := (even_iff_exists_two_nsmul _).symm.not.trans not_even_iff diff --git a/Mathlib/Algebra/Ring/Parity.lean b/Mathlib/Algebra/Ring/Parity.lean index 7b3959696657a..f74ff162454d5 100644 --- a/Mathlib/Algebra/Ring/Parity.lean +++ b/Mathlib/Algebra/Ring/Parity.lean @@ -201,7 +201,7 @@ lemma odd_iff : Odd n ↔ n % 2 = 1 := instance : DecidablePred (Odd : ℕ → Prop) := fun _ ↦ decidable_of_iff _ odd_iff.symm -lemma not_odd_iff : ¬Odd n ↔ n % 2 = 0 := by rw [odd_iff, mod_two_ne_one] +lemma not_odd_iff : ¬Odd n ↔ n % 2 = 0 := by rw [odd_iff, mod_two_not_eq_one] @[simp] lemma not_odd_iff_even : ¬Odd n ↔ Even n := by rw [not_odd_iff, even_iff] @[simp] lemma not_even_iff_odd : ¬Even n ↔ Odd n := by rw [not_even_iff, odd_iff] diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index 20c3db731c333..79583654cc19e 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -968,12 +968,15 @@ lemma set_induction {S : Set ℕ} (hb : 0 ∈ S) (h_ind : ∀ k : ℕ, k ∈ S attribute [simp] Nat.dvd_zero -@[simp] lemma mod_two_ne_one : ¬n % 2 = 1 ↔ n % 2 = 0 := by +@[simp] lemma mod_two_not_eq_one : ¬n % 2 = 1 ↔ n % 2 = 0 := by cases mod_two_eq_zero_or_one n <;> simp [*] -@[simp] lemma mod_two_ne_zero : ¬n % 2 = 0 ↔ n % 2 = 1 := by +@[simp] lemma mod_two_not_eq_zero : ¬n % 2 = 0 ↔ n % 2 = 1 := by cases mod_two_eq_zero_or_one n <;> simp [*] +lemma mod_two_ne_one : n % 2 ≠ 1 ↔ n % 2 = 0 := mod_two_not_eq_one +lemma mod_two_ne_zero : n % 2 ≠ 0 ↔ n % 2 = 1 := mod_two_not_eq_zero + @[deprecated mod_mul_right_div_self (since := "2024-05-29")] lemma div_mod_eq_mod_mul_div (a b c : ℕ) : a / b % c = a % (b * c) / b := (mod_mul_right_div_self a b c).symm From 6e1b417164a9b2dba4c7f38925e8b4ba4c22e7f1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Nov 2024 21:33:59 +0000 Subject: [PATCH 305/829] chore(Pointwise): rename `image_inv` to `image_inv_eq_inv` (#19480) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This is to make space for `f '' s⁻¹ = (f '' s)⁻¹`. Also rename `smul_card_le` to `card_smul_le`. From LeanCamCombi --- .../Algebra/Group/Pointwise/Finset/Basic.lean | 16 ++++++++----- .../Algebra/Group/Pointwise/Set/Basic.lean | 24 ++++++++++++------- Mathlib/Algebra/Group/Pointwise/Set/Card.lean | 4 ++-- .../Algebra/Group/Pointwise/Set/Finite.lean | 2 +- .../GroupWithZero/Pointwise/Finset.lean | 4 ++-- .../Group/Pointwise/CompleteLattice.lean | 8 +++---- Mathlib/Algebra/Ring/Pointwise/Finset.lean | 4 ++-- Mathlib/Algebra/Ring/Pointwise/Set.lean | 4 ++-- Mathlib/Analysis/Convex/Hull.lean | 2 +- Mathlib/Analysis/Convex/Star.lean | 2 +- Mathlib/Analysis/Convex/Topology.lean | 2 +- Mathlib/Analysis/Normed/Group/Pointwise.lean | 4 ++-- Mathlib/Analysis/Normed/Group/Quotient.lean | 2 +- Mathlib/Data/Real/Cardinality.lean | 2 +- Mathlib/Data/Set/Pointwise/SMul.lean | 8 +++---- Mathlib/MeasureTheory/Group/Arithmetic.lean | 2 +- Mathlib/MeasureTheory/Group/Prod.lean | 6 ++--- .../MeasureTheory/Measure/Haar/OfBasis.lean | 2 +- Mathlib/Topology/Algebra/Field.lean | 2 +- Mathlib/Topology/Algebra/Group/Basic.lean | 2 +- 20 files changed, 56 insertions(+), 46 deletions(-) diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index 4330e3cf9abae..ac78f4f5b38b5 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -198,9 +198,7 @@ scoped[Pointwise] attribute [instance] Finset.inv Finset.neg theorem inv_def : s⁻¹ = s.image fun x => x⁻¹ := rfl -@[to_additive] -theorem image_inv : (s.image fun x => x⁻¹) = s⁻¹ := - rfl +@[to_additive] lemma image_inv_eq_inv (s : Finset α) : s.image (·⁻¹) = s⁻¹ := rfl @[to_additive] theorem mem_inv {x : α} : x ∈ s⁻¹ ↔ ∃ y ∈ s, y⁻¹ = x := @@ -279,7 +277,7 @@ variable [DecidableEq α] [InvolutiveInv α] {s : Finset α} {a : α} lemma mem_inv' : a ∈ s⁻¹ ↔ a⁻¹ ∈ s := by simp [mem_inv, inv_eq_iff_eq_inv] @[to_additive (attr := simp, norm_cast)] -theorem coe_inv (s : Finset α) : ↑s⁻¹ = (s : Set α)⁻¹ := coe_image.trans Set.image_inv +theorem coe_inv (s : Finset α) : ↑s⁻¹ = (s : Set α)⁻¹ := coe_image.trans Set.image_inv_eq_inv @[to_additive (attr := simp)] theorem card_inv (s : Finset α) : s⁻¹.card = s.card := card_image_of_injective _ inv_injective @@ -322,7 +320,10 @@ lemma coe_smul (s : Finset α) (t : Finset β) : ↑(s • t) = (s : Set α) • @[to_additive] lemma smul_mem_smul : a ∈ s → b ∈ t → a • b ∈ s • t := mem_image₂_of_mem -@[to_additive] lemma smul_card_le : (s • t).card ≤ s.card • t.card := card_image₂_le .. +@[to_additive] lemma card_smul_le : #(s • t) ≤ #s * #t := card_image₂_le .. + +@[deprecated (since := "2024-11-19")] alias smul_card_le := card_smul_le +@[deprecated (since := "2024-11-19")] alias vadd_card_le := card_vadd_le @[to_additive (attr := simp)] lemma empty_smul (t : Finset β) : (∅ : Finset α) • t = ∅ := image₂_empty_left @@ -1173,6 +1174,9 @@ theorem image_mul_left' : theorem image_mul_right' : image (· * b⁻¹) t = preimage t (· * b) (mul_left_injective _).injOn := by simp +@[to_additive] +lemma image_inv (f : F) (s : Finset α) : s⁻¹.image f = (s.image f)⁻¹ := image_comm (map_inv _) + theorem image_div : (s / t).image (f : α → β) = s.image f / t.image f := image_image₂_distrib <| map_div f @@ -1793,4 +1797,4 @@ instance Nat.decidablePred_mem_vadd_set {s : Set ℕ} [DecidablePred (· ∈ s)] fun n ↦ decidable_of_iff' (a ≤ n ∧ n - a ∈ s) <| by simp only [Set.mem_vadd_set, vadd_eq_add]; aesop -set_option linter.style.longFile 1800 +set_option linter.style.longFile 2000 diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean index 187539c08b148..417727b33b985 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean @@ -141,10 +141,10 @@ end One section Inv /-- The pointwise inversion of set `s⁻¹` is defined as `{x | x⁻¹ ∈ s}` in locale `Pointwise`. It is -equal to `{x⁻¹ | x ∈ s}`, see `Set.image_inv`. -/ +equal to `{x⁻¹ | x ∈ s}`, see `Set.image_inv_eq_inv`. -/ @[to_additive "The pointwise negation of set `-s` is defined as `{x | -x ∈ s}` in locale `Pointwise`. - It is equal to `{-x | x ∈ s}`, see `Set.image_neg`."] + It is equal to `{-x | x ∈ s}`, see `Set.image_neg_eq_neg`."] protected def inv [Inv α] : Inv (Set α) := ⟨preimage Inv.inv⟩ @@ -218,12 +218,12 @@ theorem Nonempty.inv (h : s.Nonempty) : s⁻¹.Nonempty := nonempty_inv.2 h @[to_additive (attr := simp)] -theorem image_inv : Inv.inv '' s = s⁻¹ := +theorem image_inv_eq_inv : (·⁻¹) '' s = s⁻¹ := congr_fun (image_eq_preimage_of_inverse inv_involutive.leftInverse inv_involutive.rightInverse) _ @[to_additive (attr := simp)] theorem inv_eq_empty : s⁻¹ = ∅ ↔ s = ∅ := by - rw [← image_inv, image_eq_empty] + rw [← image_inv_eq_inv, image_eq_empty] @[to_additive (attr := simp)] noncomputable instance involutiveInv : InvolutiveInv (Set α) where @@ -238,7 +238,8 @@ theorem inv_subset_inv : s⁻¹ ⊆ t⁻¹ ↔ s ⊆ t := theorem inv_subset : s⁻¹ ⊆ t ↔ s ⊆ t⁻¹ := by rw [← inv_subset_inv, inv_inv] @[to_additive (attr := simp)] -theorem inv_singleton (a : α) : ({a} : Set α)⁻¹ = {a⁻¹} := by rw [← image_inv, image_singleton] +theorem inv_singleton (a : α) : ({a} : Set α)⁻¹ = {a⁻¹} := by + rw [← image_inv_eq_inv, image_singleton] @[to_additive (attr := simp)] theorem inv_insert (a : α) (s : Set α) : (insert a s)⁻¹ = insert a⁻¹ s⁻¹ := by @@ -246,14 +247,14 @@ 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] + rw [← image_inv_eq_inv] exact (range_comp ..).symm open MulOpposite @[to_additive] theorem image_op_inv : op '' s⁻¹ = (op '' s)⁻¹ := by - simp_rw [← image_inv, Function.Semiconj.set_image op_inv s] + simp_rw [← image_inv_eq_inv, Function.Semiconj.set_image op_inv s] end InvolutiveInv @@ -1171,13 +1172,13 @@ protected theorem mul_eq_one_iff : s * t = 1 ↔ ∃ a b, s = {a} ∧ t = {b} protected noncomputable def divisionMonoid : DivisionMonoid (Set α) := { Set.monoid, Set.involutiveInv, Set.div, @Set.ZPow α _ _ _ with mul_inv_rev := fun s t => by - simp_rw [← image_inv] + simp_rw [← image_inv_eq_inv] exact image_image2_antidistrib mul_inv_rev inv_eq_of_mul := fun s t h => by obtain ⟨a, b, rfl, rfl, hab⟩ := Set.mul_eq_one_iff.1 h rw [inv_singleton, inv_eq_of_mul_eq_one_right hab] div_eq_mul_inv := fun s t => by - rw [← image_id (s / t), ← image_inv] + rw [← image_id (s / t), ← image_inv_eq_inv] exact image_image2_distrib_right div_eq_mul_inv } scoped[Pointwise] attribute [instance] Set.divisionMonoid Set.subtractionMonoid @@ -1297,6 +1298,11 @@ 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 ..⟩ +@[to_additive] +lemma image_inv [DivisionMonoid β] [FunLike F α β] [MonoidHomClass F α β] (f : F) (s : Set α) : + f '' s⁻¹ = (f '' s)⁻¹ := by + rw [← image_inv_eq_inv, ← image_inv_eq_inv]; exact image_comm (map_inv _) + end Group section Mul diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean index d31f9773753a2..812782481817b 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean @@ -48,11 +48,11 @@ variable [InvolutiveInv G] @[to_additive (attr := simp)] lemma _root_.Cardinal.mk_inv (s : Set G) : #↥(s⁻¹) = #s := by - rw [← image_inv, Cardinal.mk_image_eq_of_injOn _ _ inv_injective.injOn] + rw [← image_inv_eq_inv, Cardinal.mk_image_eq_of_injOn _ _ inv_injective.injOn] @[to_additive (attr := simp)] lemma natCard_inv (s : Set G) : Nat.card ↥(s⁻¹) = Nat.card s := by - rw [← image_inv, Nat.card_image_of_injective inv_injective] + rw [← image_inv_eq_inv, Nat.card_image_of_injective inv_injective] @[to_additive] alias card_inv := natCard_inv diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Finite.lean b/Mathlib/Algebra/Group/Pointwise/Set/Finite.lean index 3b1d30164c535..77adb76f64340 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Finite.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Finite.lean @@ -112,7 +112,7 @@ 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] + rw [← image_inv_eq_inv, finite_image_iff inv_injective.injOn] @[to_additive (attr := simp)] lemma infinite_inv : s⁻¹.Infinite ↔ s.Infinite := finite_inv.not diff --git a/Mathlib/Algebra/GroupWithZero/Pointwise/Finset.lean b/Mathlib/Algebra/GroupWithZero/Pointwise/Finset.lean index a49b451142479..c112e0a96c541 100644 --- a/Mathlib/Algebra/GroupWithZero/Pointwise/Finset.lean +++ b/Mathlib/Algebra/GroupWithZero/Pointwise/Finset.lean @@ -199,11 +199,11 @@ variable [Monoid α] [AddGroup β] [DistribMulAction α β] @[simp] lemma smul_finset_neg (a : α) (t : Finset β) : a • -t = -(a • t) := by - simp only [← image_smul, ← image_neg, Function.comp_def, image_image, smul_neg] + simp only [← image_smul, ← image_neg_eq_neg, Function.comp_def, image_image, smul_neg] @[simp] protected lemma smul_neg (s : Finset α) (t : Finset β) : s • -t = -(s • t) := by - simp_rw [← image_neg]; exact image_image₂_right_comm smul_neg + simp_rw [← image_neg_eq_neg]; exact image_image₂_right_comm smul_neg end Monoid end Finset diff --git a/Mathlib/Algebra/Order/Group/Pointwise/CompleteLattice.lean b/Mathlib/Algebra/Order/Group/Pointwise/CompleteLattice.lean index 4a98298c00046..bdd12efab2f9b 100644 --- a/Mathlib/Algebra/Order/Group/Pointwise/CompleteLattice.lean +++ b/Mathlib/Algebra/Order/Group/Pointwise/CompleteLattice.lean @@ -39,12 +39,12 @@ variable [Group M] [MulLeftMono M] [MulRightMono M] @[to_additive] lemma csSup_inv (hs₀ : s.Nonempty) (hs₁ : BddBelow s) : sSup s⁻¹ = (sInf s)⁻¹ := by - rw [← image_inv] + rw [← image_inv_eq_inv] exact ((OrderIso.inv _).map_csInf' hs₀ hs₁).symm @[to_additive] lemma csInf_inv (hs₀ : s.Nonempty) (hs₁ : BddAbove s) : sInf s⁻¹ = (sSup s)⁻¹ := by - rw [← image_inv] + rw [← image_inv_eq_inv] exact ((OrderIso.inv _).map_csSup' hs₀ hs₁).symm @[to_additive] @@ -89,12 +89,12 @@ variable [Group M] [MulLeftMono M] [MulRightMono M] @[to_additive] lemma sSup_inv (s : Set M) : sSup s⁻¹ = (sInf s)⁻¹ := by - rw [← image_inv, sSup_image] + rw [← image_inv_eq_inv, sSup_image] exact ((OrderIso.inv M).map_sInf _).symm @[to_additive] lemma sInf_inv (s : Set M) : sInf s⁻¹ = (sSup s)⁻¹ := by - rw [← image_inv, sInf_image] + rw [← image_inv_eq_inv, sInf_image] exact ((OrderIso.inv M).map_sSup _).symm @[to_additive] diff --git a/Mathlib/Algebra/Ring/Pointwise/Finset.lean b/Mathlib/Algebra/Ring/Pointwise/Finset.lean index bd2a3b4e8a7f9..f62ea470d74ea 100644 --- a/Mathlib/Algebra/Ring/Pointwise/Finset.lean +++ b/Mathlib/Algebra/Ring/Pointwise/Finset.lean @@ -58,11 +58,11 @@ variable [Ring α] [AddCommGroup β] [Module α β] [DecidableEq β] {s : Finset @[simp] lemma neg_smul_finset : -a • t = -(a • t) := by - simp only [← image_smul, ← image_neg, image_image, neg_smul, Function.comp_def] + simp only [← image_smul, ← image_neg_eq_neg, image_image, neg_smul, Function.comp_def] @[simp] protected lemma neg_smul [DecidableEq α] : -s • t = -(s • t) := by - simp_rw [← image_neg] + simp_rw [← image_neg_eq_neg] exact image₂_image_left_comm neg_smul end Ring diff --git a/Mathlib/Algebra/Ring/Pointwise/Set.lean b/Mathlib/Algebra/Ring/Pointwise/Set.lean index 4c844172386c8..4c75fc48d5209 100644 --- a/Mathlib/Algebra/Ring/Pointwise/Set.lean +++ b/Mathlib/Algebra/Ring/Pointwise/Set.lean @@ -29,8 +29,8 @@ namespace Set /-- `Set α` has distributive negation if `α` has. -/ protected noncomputable def hasDistribNeg [Mul α] [HasDistribNeg α] : HasDistribNeg (Set α) where __ := Set.involutiveNeg - neg_mul _ _ := by simp_rw [← image_neg]; exact image2_image_left_comm neg_mul - mul_neg _ _ := by simp_rw [← image_neg]; exact image_image2_right_comm mul_neg + neg_mul _ _ := by simp_rw [← image_neg_eq_neg]; exact image2_image_left_comm neg_mul + mul_neg _ _ := by simp_rw [← image_neg_eq_neg]; exact image_image2_right_comm mul_neg scoped[Pointwise] attribute [instance] Set.hasDistribNeg diff --git a/Mathlib/Analysis/Convex/Hull.lean b/Mathlib/Analysis/Convex/Hull.lean index 846de50849e76..d627459e1acc8 100644 --- a/Mathlib/Analysis/Convex/Hull.lean +++ b/Mathlib/Analysis/Convex/Hull.lean @@ -194,7 +194,7 @@ theorem affineSpan_convexHull (s : Set E) : affineSpan 𝕜 (convexHull 𝕜 s) exact convexHull_subset_affineSpan s theorem convexHull_neg (s : Set E) : convexHull 𝕜 (-s) = -convexHull 𝕜 s := by - simp_rw [← image_neg] + simp_rw [← image_neg_eq_neg] exact AffineMap.image_convexHull (-1) _ |>.symm lemma convexHull_vadd (x : E) (s : Set E) : convexHull 𝕜 (x +ᵥ s) = x +ᵥ convexHull 𝕜 s := diff --git a/Mathlib/Analysis/Convex/Star.lean b/Mathlib/Analysis/Convex/Star.lean index ad0ee8acf1ca5..dae3f3bd26901 100644 --- a/Mathlib/Analysis/Convex/Star.lean +++ b/Mathlib/Analysis/Convex/Star.lean @@ -327,7 +327,7 @@ theorem StarConvex.affine_image (f : E →ᵃ[𝕜] F) {s : Set E} (hs : StarCon rw [Convex.combo_affine_apply hab, hy'f] theorem StarConvex.neg (hs : StarConvex 𝕜 x s) : StarConvex 𝕜 (-x) (-s) := by - rw [← image_neg] + rw [← image_neg_eq_neg] exact hs.is_linear_image IsLinearMap.isLinearMap_neg theorem StarConvex.sub (hs : StarConvex 𝕜 x s) (ht : StarConvex 𝕜 y t) : diff --git a/Mathlib/Analysis/Convex/Topology.lean b/Mathlib/Analysis/Convex/Topology.lean index 86c3bddbec48e..ac81c76f071c7 100644 --- a/Mathlib/Analysis/Convex/Topology.lean +++ b/Mathlib/Analysis/Convex/Topology.lean @@ -360,7 +360,7 @@ theorem Convex.closure_subset_image_homothety_interior_of_one_lt {s : Set E} (hs refine ⟨homothety x t⁻¹ y, hs.openSegment_interior_closure_subset_interior hx hy ?_, (AffineEquiv.homothetyUnitsMulHom x (Units.mk0 t hne)).apply_symm_apply y⟩ - rw [openSegment_eq_image_lineMap, ← inv_one, ← inv_Ioi₀ (zero_lt_one' ℝ), ← image_inv, + rw [openSegment_eq_image_lineMap, ← inv_one, ← inv_Ioi₀ (zero_lt_one' ℝ), ← image_inv_eq_inv, image_image, homothety_eq_lineMap] exact mem_image_of_mem _ ht diff --git a/Mathlib/Analysis/Normed/Group/Pointwise.lean b/Mathlib/Analysis/Normed/Group/Pointwise.lean index 8d853167e1370..d45a0621cf76b 100644 --- a/Mathlib/Analysis/Normed/Group/Pointwise.lean +++ b/Mathlib/Analysis/Normed/Group/Pointwise.lean @@ -39,7 +39,7 @@ theorem Bornology.IsBounded.of_mul (hst : IsBounded (s * t)) : IsBounded s ∨ I @[to_additive] theorem Bornology.IsBounded.inv : IsBounded s → IsBounded s⁻¹ := by - simp_rw [isBounded_iff_forall_norm_le', ← image_inv, forall_mem_image, norm_inv'] + simp_rw [isBounded_iff_forall_norm_le', ← image_inv_eq_inv, forall_mem_image, norm_inv'] exact id @[to_additive] @@ -58,7 +58,7 @@ open EMetric @[to_additive (attr := simp)] theorem infEdist_inv_inv (x : E) (s : Set E) : infEdist x⁻¹ s⁻¹ = infEdist x s := by - rw [← image_inv, infEdist_image isometry_inv] + rw [← image_inv_eq_inv, infEdist_image isometry_inv] @[to_additive] theorem infEdist_inv (x : E) (s : Set E) : infEdist x⁻¹ s = infEdist x s⁻¹ := by diff --git a/Mathlib/Analysis/Normed/Group/Quotient.lean b/Mathlib/Analysis/Normed/Group/Quotient.lean index fbdf446ba6f4a..0009c2a130e4e 100644 --- a/Mathlib/Analysis/Normed/Group/Quotient.lean +++ b/Mathlib/Analysis/Normed/Group/Quotient.lean @@ -150,7 +150,7 @@ theorem quotient_norm_mk_le' (S : AddSubgroup M) (m : M) : ‖(m : M ⧸ S)‖ /-- The norm of the image under the natural morphism to the quotient. -/ theorem quotient_norm_mk_eq (S : AddSubgroup M) (m : M) : ‖mk' S m‖ = sInf ((‖m + ·‖) '' S) := by - rw [mk'_apply, norm_mk, sInf_image', ← infDist_image isometry_neg, image_neg, + rw [mk'_apply, norm_mk, sInf_image', ← infDist_image isometry_neg, image_neg_eq_neg, neg_coe_set (H := S), infDist_eq_iInf] simp only [dist_eq_norm', sub_neg_eq_add, add_comm] diff --git a/Mathlib/Data/Real/Cardinality.lean b/Mathlib/Data/Real/Cardinality.lean index b41af8e41d56f..a4bda74e874be 100644 --- a/Mathlib/Data/Real/Cardinality.lean +++ b/Mathlib/Data/Real/Cardinality.lean @@ -252,7 +252,7 @@ theorem mk_Ioo_real {a b : ℝ} (h : a < b) : #(Ioo a b) = 𝔠 := by replace h := sub_pos_of_lt h have h2 : #(Inv.inv '' Ioo 0 (b - a)) ≤ #(Ioo 0 (b - a)) := mk_image_le refine le_trans ?_ h2 - rw [image_inv, inv_Ioo_0_left h, mk_Ioi_real] + rw [image_inv_eq_inv, inv_Ioo_0_left h, mk_Ioi_real] /-- The cardinality of the interval [a, b). -/ theorem mk_Ico_real {a b : ℝ} (h : a < b) : #(Ico a b) = 𝔠 := diff --git a/Mathlib/Data/Set/Pointwise/SMul.lean b/Mathlib/Data/Set/Pointwise/SMul.lean index e7008d4c67fbf..d37ecd85db188 100644 --- a/Mathlib/Data/Set/Pointwise/SMul.lean +++ b/Mathlib/Data/Set/Pointwise/SMul.lean @@ -521,11 +521,11 @@ variable [Monoid α] [AddGroup β] [DistribMulAction α β] (a : α) (s : Set α @[simp] theorem smul_set_neg : a • -t = -(a • t) := by - simp_rw [← image_smul, ← image_neg, image_image, smul_neg] + simp_rw [← image_smul, ← image_neg_eq_neg, image_image, smul_neg] @[simp] protected theorem smul_neg : s • -t = -(s • t) := by - simp_rw [← image_neg] + simp_rw [← image_neg_eq_neg] exact image_image2_right_comm smul_neg end Monoid @@ -546,11 +546,11 @@ variable [Ring α] [AddCommGroup β] [Module α β] (a : α) (s : Set α) (t : S @[simp] theorem neg_smul_set : -a • t = -(a • t) := by - simp_rw [← image_smul, ← image_neg, image_image, neg_smul] + simp_rw [← image_smul, ← image_neg_eq_neg, image_image, neg_smul] @[simp] protected theorem neg_smul : -s • t = -(s • t) := by - simp_rw [← image_neg] + simp_rw [← image_neg_eq_neg] exact image2_image_left_comm neg_smul end Ring diff --git a/Mathlib/MeasureTheory/Group/Arithmetic.lean b/Mathlib/MeasureTheory/Group/Arithmetic.lean index a265438e4ebd7..eaba1a41a370d 100644 --- a/Mathlib/MeasureTheory/Group/Arithmetic.lean +++ b/Mathlib/MeasureTheory/Group/Arithmetic.lean @@ -419,7 +419,7 @@ theorem MeasurableSet.inv {s : Set G} (hs : MeasurableSet s) : MeasurableSet s @[to_additive] theorem measurableEmbedding_inv [InvolutiveInv α] [MeasurableInv α] : MeasurableEmbedding (Inv.inv (α := α)) := - ⟨inv_injective, measurable_inv, fun s hs ↦ s.image_inv ▸ hs.inv⟩ + ⟨inv_injective, measurable_inv, fun s hs ↦ s.image_inv_eq_inv ▸ hs.inv⟩ end Inv diff --git a/Mathlib/MeasureTheory/Group/Prod.lean b/Mathlib/MeasureTheory/Group/Prod.lean index 8c7b2f7d77ef1..a33b49bf00b01 100644 --- a/Mathlib/MeasureTheory/Group/Prod.lean +++ b/Mathlib/MeasureTheory/Group/Prod.lean @@ -263,9 +263,9 @@ theorem ae_measure_preimage_mul_right_lt_top (hμs : μ' s ≠ ∞) : apply ae_lt_top (measurable_measure_mul_right ν' sm) have h1 := measure_mul_lintegral_eq μ' ν' sm (A⁻¹.indicator 1) (measurable_one.indicator hA.inv) rw [lintegral_indicator hA.inv] at h1 - simp_rw [Pi.one_apply, setLIntegral_one, ← image_inv, indicator_image inv_injective, image_inv, - ← indicator_mul_right _ fun x => ν' ((fun y => y * x) ⁻¹' s), Function.comp, Pi.one_apply, - mul_one] at h1 + simp_rw [Pi.one_apply, setLIntegral_one, ← image_inv_eq_inv, indicator_image inv_injective, + image_inv_eq_inv, ← indicator_mul_right _ fun x => ν' ((· * x) ⁻¹' s), Function.comp, + Pi.one_apply, mul_one] at h1 rw [← lintegral_indicator hA, ← h1] exact ENNReal.mul_ne_top hμs h3A.ne diff --git a/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean b/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean index a0db8d8e207b9..9f84eee98f624 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean @@ -116,7 +116,7 @@ theorem parallelepiped_orthonormalBasis_one_dim (b : OrthonormalBasis ι ℝ ℝ · right simp_rw [H, parallelepiped, Algebra.id.smul_eq_mul, A] simp only [F, Finset.univ_unique, Fin.default_eq_zero, mul_neg, mul_one, Finset.sum_neg_distrib, - Finset.sum_singleton, ← image_comp, Function.comp, image_neg, neg_Icc, neg_zero] + Finset.sum_singleton, ← image_comp, Function.comp, image_neg_eq_neg, neg_Icc, neg_zero] theorem parallelepiped_eq_sum_segment (v : ι → E) : parallelepiped v = ∑ i, segment ℝ 0 (v i) := by ext diff --git a/Mathlib/Topology/Algebra/Field.lean b/Mathlib/Topology/Algebra/Field.lean index cc34c6223561e..665d8fe1857c3 100644 --- a/Mathlib/Topology/Algebra/Field.lean +++ b/Mathlib/Topology/Algebra/Field.lean @@ -52,7 +52,7 @@ def Subfield.topologicalClosure (K : Subfield α) : Subfield α := rcases eq_or_ne x 0 with (rfl | h) · rwa [inv_zero] · -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11215): TODO: Lean fails to find InvMemClass instance - rw [← @inv_coe_set α (Subfield α) _ _ SubfieldClass.toInvMemClass K, ← Set.image_inv] + rw [← @inv_coe_set α (Subfield α) _ _ SubfieldClass.toInvMemClass K, ← Set.image_inv_eq_inv] exact mem_closure_image (continuousAt_inv₀ h) hx } theorem Subfield.le_topologicalClosure (s : Subfield α) : s ≤ s.topologicalClosure := diff --git a/Mathlib/Topology/Algebra/Group/Basic.lean b/Mathlib/Topology/Algebra/Group/Basic.lean index ffb45ff8d0037..4b003dcf49f11 100644 --- a/Mathlib/Topology/Algebra/Group/Basic.lean +++ b/Mathlib/Topology/Algebra/Group/Basic.lean @@ -284,7 +284,7 @@ variable [TopologicalSpace G] [InvolutiveInv G] [ContinuousInv G] {s : Set G} @[to_additive] theorem IsCompact.inv (hs : IsCompact s) : IsCompact s⁻¹ := by - rw [← image_inv] + rw [← image_inv_eq_inv] exact hs.image continuous_inv variable (G) From b50e5085f97f3825c600f239ed4a9a308572acbb Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Tue, 26 Nov 2024 03:12:44 +0000 Subject: [PATCH 306/829] chore: update Mathlib dependencies 2024-11-26 (#19497) 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 aa466b92f0770..8152b35704a8a 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "44e2d2e643fd2618b01f9a0592d7dcbd3ffa22de", + "rev": "7488499a8aad6ffada87ab6db73673d88dc04c97", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From 8471a92c0fc80a9309e82362c309441139b4eeb9 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 26 Nov 2024 04:40:53 +0000 Subject: [PATCH 307/829] chore: remove isStrictTotalOrder_of_linearOrder (#19493) This has been around since Lean3 core, but it is unused in Mathlib. Propose deletion. --- Mathlib.lean | 1 + Mathlib/Deprecated/Order.lean | 17 +++++++++++++++++ Mathlib/Order/Defs.lean | 5 ----- 3 files changed, 18 insertions(+), 5 deletions(-) create mode 100644 Mathlib/Deprecated/Order.lean diff --git a/Mathlib.lean b/Mathlib.lean index 5785a8dab0524..d0155a9daf6a8 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2903,6 +2903,7 @@ import Mathlib.Deprecated.LazyList import Mathlib.Deprecated.Logic import Mathlib.Deprecated.MinMax import Mathlib.Deprecated.NatLemmas +import Mathlib.Deprecated.Order import Mathlib.Deprecated.RelClasses import Mathlib.Deprecated.Ring import Mathlib.Deprecated.Subfield diff --git a/Mathlib/Deprecated/Order.lean b/Mathlib/Deprecated/Order.lean new file mode 100644 index 0000000000000..8e0ef3256036e --- /dev/null +++ b/Mathlib/Deprecated/Order.lean @@ -0,0 +1,17 @@ +/- +Copyright (c) 2014 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Jeremy Avigad, Floris van Doorn +-/ +import Mathlib.Order.Defs + +/-! +# Deprecated instances about order structures. +-/ + +variable {α : Type*} + +@[deprecated (since := "2024-11-26")] -- unused in Mathlib +instance isStrictTotalOrder_of_linearOrder [LinearOrder α] : IsStrictTotalOrder α (· < ·) where + irrefl := lt_irrefl + trichotomous := lt_trichotomy diff --git a/Mathlib/Order/Defs.lean b/Mathlib/Order/Defs.lean index b97cdc2c979d6..8c0f11d2df7c6 100644 --- a/Mathlib/Order/Defs.lean +++ b/Mathlib/Order/Defs.lean @@ -488,11 +488,6 @@ instance (priority := 900) (a b : α) : Decidable (a = b) := LinearOrder.decidab lemma eq_or_lt_of_not_lt (h : ¬a < b) : a = b ∨ b < a := if h₁ : a = b then Or.inl h₁ else Or.inr (lt_of_not_ge fun hge => h (lt_of_le_of_ne hge h₁)) --- TODO(Leo): decide whether we should keep this instance or not -instance isStrictTotalOrder_of_linearOrder : IsStrictTotalOrder α (· < ·) where - irrefl := lt_irrefl - trichotomous := lt_trichotomy - /-- Perform a case-split on the ordering of `x` and `y` in a decidable linear order. -/ def ltByCases (x y : α) {P : Sort*} (h₁ : x < y → P) (h₂ : x = y → P) (h₃ : y < x → P) : P := if h : x < y then h₁ h From a8f606300eaf74a2876dc206d2f030010a5b0463 Mon Sep 17 00:00:00 2001 From: Nick Ward <102917377+gio256@users.noreply.github.com> Date: Tue, 26 Nov 2024 05:23:35 +0000 Subject: [PATCH 308/829] chore(AlgebraicTopology): move quasicategory instances to Quasicategory dir (#19472) We add the file `Mathlib/AlgebraicTopology/Quasicategory/StrictSegal.lean` and move into it the proof that any strict Segal simplicial set is a quasicategory (from `Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean`). We also add the file `Mathlib/AlgebraicTopology/Quasicategory/Nerve.lean` and move into it the proof that the nerve of a category is a quasicategory (again from `Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean`). --- Mathlib.lean | 2 + .../Quasicategory/Basic.lean | 2 +- .../Quasicategory/Nerve.lean | 31 ++++++ .../Quasicategory/StrictSegal.lean | 102 ++++++++++++++++++ .../SimplicialSet/StrictSegal.lean | 80 -------------- 5 files changed, 136 insertions(+), 81 deletions(-) create mode 100644 Mathlib/AlgebraicTopology/Quasicategory/Nerve.lean create mode 100644 Mathlib/AlgebraicTopology/Quasicategory/StrictSegal.lean diff --git a/Mathlib.lean b/Mathlib.lean index d0155a9daf6a8..963adc4c3dcd2 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1035,6 +1035,8 @@ import Mathlib.AlgebraicTopology.FundamentalGroupoid.Product import Mathlib.AlgebraicTopology.FundamentalGroupoid.SimplyConnected import Mathlib.AlgebraicTopology.MooreComplex import Mathlib.AlgebraicTopology.Quasicategory.Basic +import Mathlib.AlgebraicTopology.Quasicategory.Nerve +import Mathlib.AlgebraicTopology.Quasicategory.StrictSegal import Mathlib.AlgebraicTopology.SimplexCategory import Mathlib.AlgebraicTopology.SimplicialCategory.Basic import Mathlib.AlgebraicTopology.SimplicialCategory.SimplicialObject diff --git a/Mathlib/AlgebraicTopology/Quasicategory/Basic.lean b/Mathlib/AlgebraicTopology/Quasicategory/Basic.lean index 55b27d884364c..c322b54925ac3 100644 --- a/Mathlib/AlgebraicTopology/Quasicategory/Basic.lean +++ b/Mathlib/AlgebraicTopology/Quasicategory/Basic.lean @@ -13,7 +13,7 @@ In this file we define quasicategories, a common model of infinity categories. We show that every Kan complex is a quasicategory. -In `Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean` +In `Mathlib/AlgebraicTopology/Quasicategory/Nerve.lean`, we show that the nerve of a category is a quasicategory. ## TODO diff --git a/Mathlib/AlgebraicTopology/Quasicategory/Nerve.lean b/Mathlib/AlgebraicTopology/Quasicategory/Nerve.lean new file mode 100644 index 0000000000000..22963dfe7b25f --- /dev/null +++ b/Mathlib/AlgebraicTopology/Quasicategory/Nerve.lean @@ -0,0 +1,31 @@ +/- +Copyright (c) 2024 Nick Ward. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin, Emily Riehl, Nick Ward +-/ +import Mathlib.AlgebraicTopology.Quasicategory.StrictSegal + +/-! +# The nerve of a category is a quasicategory + +In `AlgebraicTopology.Quasicategory.StrictSegal`, we show that any +strict Segal simplicial set is a quasicategory. +In `AlgebraicTopology.SimplicialSet.StrictSegal`, we show that the nerve of a +category satisfies the strict Segal condition. + +In this file, we prove as a direct consequence that the nerve of a category is +a quasicategory. +-/ + +universe v u + +open CategoryTheory SSet + +namespace Nerve + +/-- By virtue of satisfying the `StrictSegal` condition, the nerve of a +category is a `Quasicategory`. -/ +instance quasicategory {C : Type u} [Category.{v} C] : + Quasicategory (nerve C) := inferInstance + +end Nerve diff --git a/Mathlib/AlgebraicTopology/Quasicategory/StrictSegal.lean b/Mathlib/AlgebraicTopology/Quasicategory/StrictSegal.lean new file mode 100644 index 0000000000000..2cecadb7f1d91 --- /dev/null +++ b/Mathlib/AlgebraicTopology/Quasicategory/StrictSegal.lean @@ -0,0 +1,102 @@ +/- +Copyright (c) 2024 Nick Ward. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin, Emily Riehl, Nick Ward +-/ +import Mathlib.AlgebraicTopology.Quasicategory.Basic +import Mathlib.AlgebraicTopology.SimplicialSet.StrictSegal + +/-! +# Strict Segal simplicial sets are quasicategories + +In `AlgebraicTopology.SimplicialSet.StrictSegal`, we define the strict Segal +condition on a simplicial set `X`. We say that `X` is strict Segal if its +simplices are uniquely determined by their spine. + +In this file, we prove that any simplicial set satisfying the strict Segal +condition is a quasicategory. +-/ + +universe u + +open CategoryTheory +open Simplicial SimplicialObject SimplexCategory + +namespace SSet.StrictSegal + +/-- Any `StrictSegal` simplicial set is a `Quasicategory`. -/ +instance quasicategory {X : SSet.{u}} [StrictSegal X] : Quasicategory X := by + apply quasicategory_of_filler X + intro n i σ₀ h₀ hₙ + use spineToSimplex <| Path.map (horn.spineId i h₀ hₙ) σ₀ + intro j hj + apply spineInjective + ext k + dsimp only [spineEquiv, spine_arrow, Function.comp_apply, Equiv.coe_fn_mk] + rw [← types_comp_apply (σ₀.app _) (X.map _), ← σ₀.naturality] + let ksucc := k.succ.castSucc + obtain hlt | hgt | heq : ksucc < j ∨ j < ksucc ∨ j = ksucc := by omega + · rw [← spine_arrow, spine_δ_arrow_lt _ hlt] + dsimp only [Path.map, spine_arrow, Fin.coe_eq_castSucc] + apply congr_arg + simp only [horn, horn.spineId, standardSimplex, uliftFunctor, Functor.comp_obj, + yoneda_obj_obj, whiskering_obj_obj_map, uliftFunctor_map, yoneda_obj_map, + standardSimplex.objEquiv, Equiv.ulift, Equiv.coe_fn_symm_mk, + Quiver.Hom.unop_op, horn.face_coe, Subtype.mk.injEq] + rw [mkOfSucc_δ_lt hlt] + rfl + · rw [← spine_arrow, spine_δ_arrow_gt _ hgt] + dsimp only [Path.map, spine_arrow, Fin.coe_eq_castSucc] + apply congr_arg + simp only [horn, horn.spineId, standardSimplex, uliftFunctor, Functor.comp_obj, + yoneda_obj_obj, whiskering_obj_obj_map, uliftFunctor_map, yoneda_obj_map, + standardSimplex.objEquiv, Equiv.ulift, Equiv.coe_fn_symm_mk, + Quiver.Hom.unop_op, horn.face_coe, Subtype.mk.injEq] + rw [mkOfSucc_δ_gt hgt] + rfl + · /- The only inner horn of `Δ[2]` does not contain the diagonal edge. -/ + have hn0 : n ≠ 0 := by + rintro rfl + obtain rfl : k = 0 := by omega + fin_cases i <;> contradiction + /- We construct the triangle in the standard simplex as a 2-simplex in + the horn. While the triangle is not contained in the inner horn `Λ[2, 1]`, + we can inhabit `Λ[n + 2, i] _[2]` by induction on `n`. -/ + let triangle : Λ[n + 2, i] _[2] := by + cases n with + | zero => contradiction + | succ _ => exact horn.primitiveTriangle i h₀ hₙ k (by omega) + /- The interval spanning from `k` to `k + 2` is equivalently the spine + of the triangle with vertices `k`, `k + 1`, and `k + 2`. -/ + have hi : ((horn.spineId i h₀ hₙ).map σ₀).interval k 2 (by omega) = + X.spine 2 (σ₀.app _ triangle) := by + ext m + dsimp [spine_arrow, Path.interval, Path.map] + rw [← types_comp_apply (σ₀.app _) (X.map _), ← σ₀.naturality] + apply congr_arg + simp only [horn, standardSimplex, uliftFunctor, Functor.comp_obj, + whiskering_obj_obj_obj, yoneda_obj_obj, uliftFunctor_obj, ne_eq, + whiskering_obj_obj_map, uliftFunctor_map, yoneda_obj_map, len_mk, + Nat.reduceAdd, Quiver.Hom.unop_op] + cases n with + | zero => contradiction + | succ _ => ext x; fin_cases x <;> fin_cases m <;> rfl + rw [← spine_arrow, spine_δ_arrow_eq _ heq, hi] + simp only [spineToDiagonal, diagonal, spineToSimplex_spine] + rw [← types_comp_apply (σ₀.app _) (X.map _), ← σ₀.naturality, types_comp_apply] + apply congr_arg + simp only [horn, standardSimplex, uliftFunctor, Functor.comp_obj, + whiskering_obj_obj_obj, yoneda_obj_obj, uliftFunctor_obj, + uliftFunctor_map, whiskering_obj_obj_map, yoneda_obj_map, horn.face_coe, + len_mk, Nat.reduceAdd, Quiver.Hom.unop_op, Subtype.mk.injEq, ULift.up_inj] + ext z + cases n with + | zero => contradiction + | succ _ => + fin_cases z <;> + · simp only [standardSimplex.objEquiv, uliftFunctor_map, yoneda_obj_map, + Quiver.Hom.unop_op, Equiv.ulift_symm_down] + rw [mkOfSucc_δ_eq heq] + rfl + +end SSet.StrictSegal diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean b/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean index cef9a156cab5a..71e41b710ae25 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean @@ -3,7 +3,6 @@ Copyright (c) 2024 Emily Riehl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Emily Riehl, Joël Riou, Johan Commelin, Nick Ward -/ -import Mathlib.AlgebraicTopology.Quasicategory.Basic import Mathlib.AlgebraicTopology.SimplicialSet.Nerve import Mathlib.AlgebraicTopology.SimplicialSet.Path import Mathlib.CategoryTheory.Functor.KanExtension.Adjunction @@ -153,81 +152,6 @@ lemma spine_δ_arrow_eq (f : Path X (n + 1)) {i : Fin n} {j : Fin (n + 2)} rw [← FunctorToTypes.map_comp_apply, ← op_comp] rw [mkOfSucc_δ_eq h, spineToSimplex_edge] -/-- Any `StrictSegal` simplicial set is a `Quasicategory`. -/ -instance : Quasicategory X := by - apply quasicategory_of_filler X - intro n i σ₀ h₀ hₙ - use spineToSimplex <| Path.map (horn.spineId i h₀ hₙ) σ₀ - intro j hj - apply spineInjective - ext k - · dsimp only [spineEquiv, spine_arrow, Function.comp_apply, Equiv.coe_fn_mk] - rw [← types_comp_apply (σ₀.app _) (X.map _), ← σ₀.naturality] - let ksucc := k.succ.castSucc - obtain hlt | hgt | heq : ksucc < j ∨ j < ksucc ∨ j = ksucc := by omega - · rw [← spine_arrow, spine_δ_arrow_lt _ hlt] - dsimp only [Path.map, spine_arrow, Fin.coe_eq_castSucc] - apply congr_arg - simp only [horn, horn.spineId, standardSimplex, uliftFunctor, Functor.comp_obj, - yoneda_obj_obj, whiskering_obj_obj_map, uliftFunctor_map, yoneda_obj_map, - standardSimplex.objEquiv, Equiv.ulift, Equiv.coe_fn_symm_mk, - Quiver.Hom.unop_op, horn.face_coe, Subtype.mk.injEq] - rw [mkOfSucc_δ_lt hlt] - rfl - · rw [← spine_arrow, spine_δ_arrow_gt _ hgt] - dsimp only [Path.map, spine_arrow, Fin.coe_eq_castSucc] - apply congr_arg - simp only [horn, horn.spineId, standardSimplex, uliftFunctor, Functor.comp_obj, - yoneda_obj_obj, whiskering_obj_obj_map, uliftFunctor_map, yoneda_obj_map, - standardSimplex.objEquiv, Equiv.ulift, Equiv.coe_fn_symm_mk, - Quiver.Hom.unop_op, horn.face_coe, Subtype.mk.injEq] - rw [mkOfSucc_δ_gt hgt] - rfl - · /- The only inner horn of `Δ[2]` does not contain the diagonal edge. -/ - have hn0 : n ≠ 0 := by - rintro rfl - obtain rfl : k = 0 := by omega - fin_cases i <;> contradiction - /- We construct the triangle in the standard simplex as a 2-simplex in - the horn. While the triangle is not contained in the inner horn `Λ[2, 1]`, - we can inhabit `Λ[n + 2, i] _[2]` by induction on `n`. -/ - let triangle : Λ[n + 2, i] _[2] := by - cases n with - | zero => contradiction - | succ _ => exact horn.primitiveTriangle i h₀ hₙ k (by omega) - /- The interval spanning from `k` to `k + 2` is equivalently the spine - of the triangle with vertices `k`, `k + 1`, and `k + 2`. -/ - have hi : ((horn.spineId i h₀ hₙ).map σ₀).interval k 2 (by omega) = - X.spine 2 (σ₀.app _ triangle) := by - ext m - dsimp [spine_arrow, Path.interval, Path.map] - rw [← types_comp_apply (σ₀.app _) (X.map _), ← σ₀.naturality] - apply congr_arg - simp only [horn, standardSimplex, uliftFunctor, Functor.comp_obj, - whiskering_obj_obj_obj, yoneda_obj_obj, uliftFunctor_obj, ne_eq, - whiskering_obj_obj_map, uliftFunctor_map, yoneda_obj_map, len_mk, - Nat.reduceAdd, Quiver.Hom.unop_op] - cases n with - | zero => contradiction - | succ _ => ext x; fin_cases x <;> fin_cases m <;> rfl - rw [← spine_arrow, spine_δ_arrow_eq _ heq, hi] - simp only [spineToDiagonal, diagonal, spineToSimplex_spine] - rw [← types_comp_apply (σ₀.app _) (X.map _), ← σ₀.naturality, types_comp_apply] - apply congr_arg - simp only [horn, standardSimplex, uliftFunctor, Functor.comp_obj, - whiskering_obj_obj_obj, yoneda_obj_obj, uliftFunctor_obj, - uliftFunctor_map, whiskering_obj_obj_map, yoneda_obj_map, horn.face_coe, - len_mk, Nat.reduceAdd, Quiver.Hom.unop_op, Subtype.mk.injEq, ULift.up_inj] - ext z - cases n with - | zero => contradiction - | succ _ => - fin_cases z <;> - · simp only [standardSimplex.objEquiv, uliftFunctor_map, yoneda_obj_map, - Quiver.Hom.unop_op, Equiv.ulift_symm_down] - rw [mkOfSucc_δ_eq heq] - rfl - end StrictSegal end SSet @@ -260,8 +184,4 @@ noncomputable instance strictSegal (C : Type u) [Category.{v} C] : StrictSegal ( · intro i hi apply ComposableArrows.mkOfObjOfMapSucc_map_succ -/-- By virtue of satisfying the `StrictSegal` condition, the nerve of a -category is a `Quasicategory`. -/ -instance : Quasicategory (nerve C) := inferInstance - end Nerve From c912bed75af20c28c1a7c67461a3bc31a0a8711a Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Tue, 26 Nov 2024 05:23:36 +0000 Subject: [PATCH 309/829] chore(NumberTheory/LSeries/PrimesInAP): fix name (Lfunction -> LFunction) (#19474) --- Mathlib/NumberTheory/LSeries/PrimesInAP.lean | 56 ++++++++++---------- 1 file changed, 28 insertions(+), 28 deletions(-) diff --git a/Mathlib/NumberTheory/LSeries/PrimesInAP.lean b/Mathlib/NumberTheory/LSeries/PrimesInAP.lean index 3cd0c083dd5d6..840e9ccbfa672 100644 --- a/Mathlib/NumberTheory/LSeries/PrimesInAP.lean +++ b/Mathlib/NumberTheory/LSeries/PrimesInAP.lean @@ -27,14 +27,14 @@ The main steps of the proof are as follows. agrees (on `re s > 1`) with the corresponding linear combination of negative logarithmic derivatives of Dirichlet L-functions. See `ArithmeticFunction.vonMangoldt.LSeries_residueClass_eq`. -4. Define an auxiliary function `ArithmeticFunction.vonMangoldt.LfunctionResidueClassAux a` that is +4. Define an auxiliary function `ArithmeticFunction.vonMangoldt.LFunctionResidueClassAux a` that is this linear combination of negative logarithmic derivatives of L-functions minus `(q.totient)⁻¹/(s-1)`, which cancels the pole at `s = 1`. - See `ArithmeticFunction.vonMangoldt.eqOn_LfunctionResidueClassAux` for the statement + See `ArithmeticFunction.vonMangoldt.eqOn_LFunctionResidueClassAux` for the statement that the auxiliary function agrees with the L-series of `ArithmeticFunction.vonMangoldt.residueClass` up to the term `(q.totient)⁻¹/(s-1)`. 5. Show that the auxiliary function is continuous on `re s ≥ 1`; - see `ArithmeticFunction.vonMangoldt.continuousOn_LfunctionResidueClassAux`. + see `ArithmeticFunction.vonMangoldt.continuousOn_LFunctionResidueClassAux`. This relies heavily on the non-vanishing of Dirichlet L-functions on the *closed* half-plane `re s ≥ 1` (`DirichletCharacter.LFunction_ne_zero_of_one_le_re`), which in turn can only be stated since we know that the L-series of a Dirichlet character @@ -51,7 +51,7 @@ The main steps of the proof are as follows. ## Definitions * `ArithmeticFunction.vonMangoldt.residueClass a` (see above). -* `ArithmeticFunction.vonMangoldt.continuousOn_LfunctionResidueClassAux` (see above). +* `ArithmeticFunction.vonMangoldt.continuousOn_LFunctionResidueClassAux` (see above). ## Main Result @@ -298,19 +298,19 @@ open Classical in Dirichlet's Theorem. On `re s > 1`, it agrees with the L-series of the von Mangoldt function restricted to the residue class `a : ZMod q` minus the principal part `(q.totient)⁻¹/(s-1)` of the pole at `s = 1`; -see `ArithmeticFunction.vonMangoldt.eqOn_LfunctionResidueClassAux`. -/ +see `ArithmeticFunction.vonMangoldt.eqOn_LFunctionResidueClassAux`. -/ noncomputable -abbrev LfunctionResidueClassAux (s : ℂ) : ℂ := +abbrev LFunctionResidueClassAux (s : ℂ) : ℂ := (q.totient : ℂ)⁻¹ * (-deriv (LFunctionTrivChar₁ q) s / LFunctionTrivChar₁ q s - ∑ χ ∈ ({1}ᶜ : Finset (DirichletCharacter ℂ q)), χ a⁻¹ * deriv (LFunction χ) s / LFunction χ s) /-- The auxiliary function is continuous away from the zeros of the L-functions of the Dirichlet characters mod `q` (including at `s = 1`). -/ -lemma continuousOn_LfunctionResidueClassAux' : - ContinuousOn (LfunctionResidueClassAux a) +lemma continuousOn_LFunctionResidueClassAux' : + ContinuousOn (LFunctionResidueClassAux a) {s | s = 1 ∨ ∀ χ : DirichletCharacter ℂ q, LFunction χ s ≠ 0} := by - rw [show LfunctionResidueClassAux a = fun s ↦ _ from rfl] - simp only [LfunctionResidueClassAux, sub_eq_add_neg] + rw [show LFunctionResidueClassAux a = fun s ↦ _ from rfl] + simp only [LFunctionResidueClassAux, sub_eq_add_neg] refine continuousOn_const.mul <| ContinuousOn.add ?_ ?_ · refine (continuousOn_neg_logDeriv_LFunctionTrivChar₁ q).mono fun s hs ↦ ?_ have := LFunction_ne_zero_of_one_le_re (1 : DirichletCharacter ℂ q) (s := s) @@ -328,11 +328,11 @@ lemma continuousOn_LfunctionResidueClassAux' : /-- The L-series of the von Mangoldt function restricted to the prime residue class `a` mod `q` is continuous on `re s ≥ 1` except for a simple pole at `s = 1` with residue `(q.totient)⁻¹`. -The statement as given here in terms of `ArithmeticFunction.vonMangoldt.LfunctionResidueClassAux` +The statement as given here in terms of `ArithmeticFunction.vonMangoldt.LFunctionResidueClassAux` is equivalent. -/ -lemma continuousOn_LfunctionResidueClassAux : - ContinuousOn (LfunctionResidueClassAux a) {s | 1 ≤ s.re} := by - refine (continuousOn_LfunctionResidueClassAux' a).mono fun s hs ↦ ?_ +lemma continuousOn_LFunctionResidueClassAux : + ContinuousOn (LFunctionResidueClassAux a) {s | 1 ≤ s.re} := by + refine (continuousOn_LFunctionResidueClassAux' a).mono fun s hs ↦ ?_ rcases eq_or_ne s 1 with rfl | hs₁ · simp only [ne_eq, Set.mem_setOf_eq, true_or] · simp only [ne_eq, Set.mem_setOf_eq, hs₁, false_or] @@ -345,13 +345,13 @@ open scoped LSeries.notation /-- The auxiliary function agrees on `re s > 1` with the L-series of the von Mangoldt function restricted to the residue class `a : ZMod q` minus the principal part `(q.totient)⁻¹/(s-1)` of its pole at `s = 1`. -/ -lemma eqOn_LfunctionResidueClassAux (ha : IsUnit a) : - Set.EqOn (LfunctionResidueClassAux a) +lemma eqOn_LFunctionResidueClassAux (ha : IsUnit a) : + Set.EqOn (LFunctionResidueClassAux a) (fun s ↦ L ↗(residueClass a) s - (q.totient : ℂ)⁻¹ / (s - 1)) {s | 1 < s.re} := by intro s hs replace hs := Set.mem_setOf.mp hs - simp only [LSeries_residueClass_eq ha hs, LfunctionResidueClassAux] + simp only [LSeries_residueClass_eq ha hs, LFunctionResidueClassAux] rw [neg_div, ← neg_add', mul_neg, ← neg_mul, div_eq_mul_one_div (q.totient : ℂ)⁻¹, sub_eq_add_neg, ← neg_mul, ← mul_add] congrm (_ * ?_) @@ -369,9 +369,9 @@ lemma eqOn_LfunctionResidueClassAux (ha : IsUnit a) : rw [mul_comm _ 1, mul_div_mul_right _ _ <| LFunction_ne_zero_of_one_le_re 1 (.inr hs₁) hs.le] /-- The auxiliary function takes real values for real arguments `x > 1`. -/ -lemma LfunctionResidueClassAux_real (ha : IsUnit a) {x : ℝ} (hx : 1 < x) : - LfunctionResidueClassAux a x = (LfunctionResidueClassAux a x).re := by - rw [eqOn_LfunctionResidueClassAux ha hx] +lemma LFunctionResidueClassAux_real (ha : IsUnit a) {x : ℝ} (hx : 1 < x) : + LFunctionResidueClassAux a x = (LFunctionResidueClassAux a x).re := by + rw [eqOn_LFunctionResidueClassAux ha hx] simp only [sub_re, ofReal_sub] congr 1 · rw [LSeries, re_tsum <| LSeriesSummable_of_abscissaOfAbsConv_lt_re <| @@ -394,25 +394,25 @@ lemma LSeries_residueClass_lower_bound (ha : IsUnit a) : (q.totient : ℝ)⁻¹ / (x - 1) - C ≤ ∑' n, residueClass a n / (n : ℝ) ^ x := by have H {x : ℝ} (hx : 1 < x) : ∑' n, residueClass a n / (n : ℝ) ^ x = - (LfunctionResidueClassAux a x).re + (q.totient : ℝ)⁻¹ / (x - 1) := by + (LFunctionResidueClassAux a x).re + (q.totient : ℝ)⁻¹ / (x - 1) := by refine ofReal_injective ?_ simp only [ofReal_tsum, ofReal_div, ofReal_cpow (Nat.cast_nonneg _), ofReal_natCast, ofReal_add, ofReal_inv, ofReal_sub, ofReal_one] - simp_rw [← LfunctionResidueClassAux_real ha hx, - eqOn_LfunctionResidueClassAux ha <| Set.mem_setOf.mpr (ofReal_re x ▸ hx), sub_add_cancel, + simp_rw [← LFunctionResidueClassAux_real ha hx, + eqOn_LFunctionResidueClassAux ha <| Set.mem_setOf.mpr (ofReal_re x ▸ hx), sub_add_cancel, LSeries, term] refine tsum_congr fun n ↦ ?_ split_ifs with hn · simp only [hn, residueClass_apply_zero, ofReal_zero, zero_div] · rfl - have : ContinuousOn (fun x : ℝ ↦ (LfunctionResidueClassAux a x).re) (Set.Icc 1 2) := - continuous_re.continuousOn.comp (t := Set.univ) (continuousOn_LfunctionResidueClassAux a) + have : ContinuousOn (fun x : ℝ ↦ (LFunctionResidueClassAux a x).re) (Set.Icc 1 2) := + continuous_re.continuousOn.comp (t := Set.univ) (continuousOn_LFunctionResidueClassAux a) (fun ⦃x⦄ a ↦ trivial) |>.comp continuous_ofReal.continuousOn fun x hx ↦ by simpa only [Set.mem_setOf_eq, ofReal_re] using hx.1 obtain ⟨C, hC⟩ := bddBelow_def.mp <| IsCompact.bddBelow_image isCompact_Icc this - replace hC {x : ℝ} (hx : x ∈ Set.Icc 1 2) : C ≤ (LfunctionResidueClassAux a x).re := - hC (LfunctionResidueClassAux a x).re <| - Set.mem_image_of_mem (fun x : ℝ ↦ (LfunctionResidueClassAux a x).re) hx + replace hC {x : ℝ} (hx : x ∈ Set.Icc 1 2) : C ≤ (LFunctionResidueClassAux a x).re := + hC (LFunctionResidueClassAux a x).re <| + Set.mem_image_of_mem (fun x : ℝ ↦ (LFunctionResidueClassAux a x).re) hx refine ⟨-C, fun {x} hx ↦ ?_⟩ rw [H hx.1, add_comm, sub_neg_eq_add, add_le_add_iff_left] exact hC <| Set.mem_Icc_of_Ioc hx From b58a55482b9cbe092750d326c261e04667ef0c04 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 26 Nov 2024 05:23:37 +0000 Subject: [PATCH 310/829] chore: split Order.BoundedOrder (#19488) --- .../OrderedCancelAddCommMonoidWithBounds.lean | 2 +- Mathlib.lean | 4 +- Mathlib/Algebra/Associated/Basic.lean | 2 +- .../Algebra/Order/Monoid/Canonical/Defs.lean | 2 +- .../Order/Monoid/Unbundled/TypeTags.lean | 2 +- Mathlib/Data/List/MinMax.lean | 1 + Mathlib/Data/PSigma/Order.lean | 3 +- Mathlib/Data/Prod/Lex.lean | 3 +- Mathlib/Data/Sigma/Order.lean | 3 +- .../Basic.lean} | 229 +----------------- Mathlib/Order/BoundedOrder/Lattice.lean | 165 +++++++++++++ Mathlib/Order/BoundedOrder/Monotone.lean | 124 ++++++++++ Mathlib/Order/Bounds/Basic.lean | 1 + Mathlib/Order/Disjoint.lean | 2 +- Mathlib/Order/Hom/Bounded.lean | 2 +- Mathlib/Order/Nat.lean | 2 +- Mathlib/Order/SuccPred/Limit.lean | 2 +- Mathlib/Order/WithBot.lean | 3 +- scripts/noshake.json | 1 + 19 files changed, 315 insertions(+), 238 deletions(-) rename Mathlib/Order/{BoundedOrder.lean => BoundedOrder/Basic.lean} (69%) create mode 100644 Mathlib/Order/BoundedOrder/Lattice.lean create mode 100644 Mathlib/Order/BoundedOrder/Monotone.lean diff --git a/Counterexamples/OrderedCancelAddCommMonoidWithBounds.lean b/Counterexamples/OrderedCancelAddCommMonoidWithBounds.lean index ba098aab1da69..16f749c3c471f 100644 --- a/Counterexamples/OrderedCancelAddCommMonoidWithBounds.lean +++ b/Counterexamples/OrderedCancelAddCommMonoidWithBounds.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Martin Dvorak -/ import Mathlib.Algebra.Order.Monoid.Defs -import Mathlib.Order.BoundedOrder +import Mathlib.Order.BoundedOrder.Lattice /-! # Do not combine OrderedCancelAddCommMonoid with BoundedOrder diff --git a/Mathlib.lean b/Mathlib.lean index 963adc4c3dcd2..fddf683280feb 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3888,7 +3888,9 @@ import Mathlib.Order.BooleanAlgebra import Mathlib.Order.BooleanGenerators import Mathlib.Order.Booleanisation import Mathlib.Order.Bounded -import Mathlib.Order.BoundedOrder +import Mathlib.Order.BoundedOrder.Basic +import Mathlib.Order.BoundedOrder.Lattice +import Mathlib.Order.BoundedOrder.Monotone import Mathlib.Order.Bounds.Basic import Mathlib.Order.Bounds.Defs import Mathlib.Order.Bounds.Image diff --git a/Mathlib/Algebra/Associated/Basic.lean b/Mathlib/Algebra/Associated/Basic.lean index 91d3f4ef3b56d..4484964fb4604 100644 --- a/Mathlib/Algebra/Associated/Basic.lean +++ b/Mathlib/Algebra/Associated/Basic.lean @@ -8,9 +8,9 @@ import Mathlib.Algebra.GroupWithZero.Divisibility import Mathlib.Algebra.GroupWithZero.Hom import Mathlib.Algebra.Group.Commute.Units import Mathlib.Algebra.Group.Units.Equiv -import Mathlib.Order.BoundedOrder import Mathlib.Algebra.Ring.Units import Mathlib.Algebra.Prime.Lemmas +import Mathlib.Order.BoundedOrder.Basic /-! # Associated elements. diff --git a/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean b/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean index b18b3c81d456e..a6bc882dcc6ef 100644 --- a/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean +++ b/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean @@ -7,7 +7,7 @@ import Mathlib.Algebra.Group.Units.Basic import Mathlib.Algebra.Order.Monoid.Defs import Mathlib.Algebra.Order.Monoid.Unbundled.ExistsOfLE import Mathlib.Algebra.NeZero -import Mathlib.Order.BoundedOrder +import Mathlib.Order.BoundedOrder.Basic /-! # Canonically ordered monoids diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean index 8a7c336d6901c..761ae23ae5a20 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl -/ import Mathlib.Algebra.Order.Monoid.Unbundled.ExistsOfLE -import Mathlib.Order.BoundedOrder import Mathlib.Algebra.Group.TypeTags.Basic +import Mathlib.Order.BoundedOrder.Basic /-! # Ordered monoid structures on `Multiplicative α` and `Additive α`. -/ diff --git a/Mathlib/Data/List/MinMax.lean b/Mathlib/Data/List/MinMax.lean index 0237438fa2eda..2a7dbf99b60f2 100644 --- a/Mathlib/Data/List/MinMax.lean +++ b/Mathlib/Data/List/MinMax.lean @@ -6,6 +6,7 @@ Authors: Minchao Wu, Chris Hughes, Mantas Bakšys import Mathlib.Data.List.Basic import Mathlib.Order.MinMax import Mathlib.Order.WithBot +import Mathlib.Order.BoundedOrder.Lattice /-! # Minimum and maximum of lists diff --git a/Mathlib/Data/PSigma/Order.lean b/Mathlib/Data/PSigma/Order.lean index c4337051e8a85..a91ce1e8563bc 100644 --- a/Mathlib/Data/PSigma/Order.lean +++ b/Mathlib/Data/PSigma/Order.lean @@ -4,10 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison, Minchao Wu -/ import Mathlib.Data.Sigma.Lex -import Mathlib.Order.BoundedOrder import Mathlib.Util.Notation3 import Init.NotationExtra import Mathlib.Data.Sigma.Basic +import Mathlib.Order.Lattice +import Mathlib.Order.BoundedOrder.Basic /-! # Lexicographic order on a sigma type diff --git a/Mathlib/Data/Prod/Lex.lean b/Mathlib/Data/Prod/Lex.lean index c248e43cdc322..2aafe976d991a 100644 --- a/Mathlib/Data/Prod/Lex.lean +++ b/Mathlib/Data/Prod/Lex.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison, Minchao Wu -/ import Mathlib.Data.Prod.Basic -import Mathlib.Order.BoundedOrder +import Mathlib.Order.Lattice +import Mathlib.Order.BoundedOrder.Basic /-! # Lexicographic order diff --git a/Mathlib/Data/Sigma/Order.lean b/Mathlib/Data/Sigma/Order.lean index 3afb990febe69..7e986738b2dc0 100644 --- a/Mathlib/Data/Sigma/Order.lean +++ b/Mathlib/Data/Sigma/Order.lean @@ -4,9 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Data.Sigma.Lex -import Mathlib.Order.BoundedOrder import Mathlib.Util.Notation3 import Mathlib.Data.Sigma.Basic +import Mathlib.Order.Lattice +import Mathlib.Order.BoundedOrder.Basic /-! # Orders on a sigma type diff --git a/Mathlib/Order/BoundedOrder.lean b/Mathlib/Order/BoundedOrder/Basic.lean similarity index 69% rename from Mathlib/Order/BoundedOrder.lean rename to Mathlib/Order/BoundedOrder/Basic.lean index 87bd2c47b7d17..e767d449e62e6 100644 --- a/Mathlib/Order/BoundedOrder.lean +++ b/Mathlib/Order/BoundedOrder/Basic.lean @@ -3,10 +3,11 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ -import Mathlib.Order.Lattice +import Mathlib.Order.Max import Mathlib.Order.ULift import Mathlib.Tactic.PushNeg import Mathlib.Tactic.Finiteness.Attr +import Mathlib.Util.AssertExists /-! # ⊤ and ⊥, bounded lattices and variants @@ -21,15 +22,10 @@ instances for `Prop` and `fun`. * `Order α`: Order with a top/bottom element. * `BoundedOrder α`: Order with a top and bottom element. -## Common lattices - -* Distributive lattices with a bottom element. Notated by `[DistribLattice α] [OrderBot α]` - It captures the properties of `Disjoint` that are common to `GeneralizedBooleanAlgebra` and - `DistribLattice` when `OrderBot`. -* Bounded and distributive lattice. Notated by `[DistribLattice α] [BoundedOrder α]`. - Typical examples include `Prop` and `Det α`. -/ +assert_not_exists Monotone + open Function OrderDual universe u v @@ -149,12 +145,6 @@ theorem Ne.lt_top' (h : ⊤ ≠ a) : a < ⊤ := theorem ne_top_of_le_ne_top (hb : b ≠ ⊤) (hab : a ≤ b) : a ≠ ⊤ := (hab.trans_lt hb.lt_top).ne -theorem StrictMono.apply_eq_top_iff (hf : StrictMono f) : f a = f ⊤ ↔ a = ⊤ := - ⟨fun h => not_lt_top_iff.1 fun ha => (hf ha).ne h, congr_arg _⟩ - -theorem StrictAnti.apply_eq_top_iff (hf : StrictAnti f) : f a = f ⊤ ↔ a = ⊤ := - ⟨fun h => not_lt_top_iff.1 fun ha => (hf ha).ne' h, congr_arg _⟩ - lemma top_not_mem_iff {s : Set α} : ⊤ ∉ s ↔ ∀ x ∈ s, x < ⊤ := ⟨fun h x hx ↦ Ne.lt_top (fun hx' : x = ⊤ ↦ h (hx' ▸ hx)), fun h h₀ ↦ (h ⊤ h₀).false⟩ @@ -166,14 +156,6 @@ theorem not_isMin_top : ¬IsMin (⊤ : α) := fun h => end OrderTop -theorem StrictMono.maximal_preimage_top [LinearOrder α] [Preorder β] [OrderTop β] {f : α → β} - (H : StrictMono f) {a} (h_top : f a = ⊤) (x : α) : x ≤ a := - H.maximal_of_maximal_image - (fun p => by - rw [h_top] - exact le_top) - x - theorem OrderTop.ext_top {α} {hA : PartialOrder α} (A : OrderTop α) {hB : PartialOrder α} (B : OrderTop α) (H : ∀ x y : α, (haveI := hA; x ≤ y) ↔ x ≤ y) : (@Top.top α (@OrderTop.toTop α hA.toLE A)) = (@Top.top α (@OrderTop.toTop α hB.toLE B)) := by @@ -326,12 +308,6 @@ theorem Ne.bot_lt' (h : ⊥ ≠ a) : ⊥ < a := theorem ne_bot_of_le_ne_bot (hb : b ≠ ⊥) (hab : b ≤ a) : a ≠ ⊥ := (hb.bot_lt.trans_le hab).ne' -theorem StrictMono.apply_eq_bot_iff (hf : StrictMono f) : f a = f ⊥ ↔ a = ⊥ := - hf.dual.apply_eq_top_iff - -theorem StrictAnti.apply_eq_bot_iff (hf : StrictAnti f) : f a = f ⊥ ↔ a = ⊥ := - hf.dual.apply_eq_top_iff - lemma bot_not_mem_iff {s : Set α} : ⊥ ∉ s ↔ ∀ x ∈ s, ⊥ < x := top_not_mem_iff (α := αᵒᵈ) @@ -342,14 +318,6 @@ theorem not_isMax_bot : ¬IsMax (⊥ : α) := end OrderBot -theorem StrictMono.minimal_preimage_bot [LinearOrder α] [PartialOrder β] [OrderBot β] {f : α → β} - (H : StrictMono f) {a} (h_bot : f a = ⊥) (x : α) : a ≤ x := - H.minimal_of_minimal_image - (fun p => by - rw [h_bot] - exact bot_le) - x - theorem OrderBot.ext_bot {α} {hA : PartialOrder α} (A : OrderBot α) {hB : PartialOrder α} (B : OrderBot α) (H : ∀ x y : α, (haveI := hA; x ≤ y) ↔ x ≤ y) : (@Bot.bot α (@OrderBot.toBot α hA.toLE A)) = (@Bot.bot α (@OrderBot.toBot α hB.toLE B)) := by @@ -357,65 +325,6 @@ theorem OrderBot.ext_bot {α} {hA : PartialOrder α} (A : OrderBot α) {hB : Par apply bot_unique exact @bot_le _ _ A _ -section SemilatticeSupTop - -variable [SemilatticeSup α] [OrderTop α] - --- Porting note: Not simp because simp can prove it -theorem top_sup_eq (a : α) : ⊤ ⊔ a = ⊤ := - sup_of_le_left le_top - --- Porting note: Not simp because simp can prove it -theorem sup_top_eq (a : α) : a ⊔ ⊤ = ⊤ := - sup_of_le_right le_top - -end SemilatticeSupTop - -section SemilatticeSupBot - -variable [SemilatticeSup α] [OrderBot α] {a b : α} - --- Porting note: Not simp because simp can prove it -theorem bot_sup_eq (a : α) : ⊥ ⊔ a = a := - sup_of_le_right bot_le - --- Porting note: Not simp because simp can prove it -theorem sup_bot_eq (a : α) : a ⊔ ⊥ = a := - sup_of_le_left bot_le - -@[simp] -theorem sup_eq_bot_iff : a ⊔ b = ⊥ ↔ a = ⊥ ∧ b = ⊥ := by rw [eq_bot_iff, sup_le_iff]; simp - -end SemilatticeSupBot - -section SemilatticeInfTop - -variable [SemilatticeInf α] [OrderTop α] {a b : α} - --- Porting note: Not simp because simp can prove it -lemma top_inf_eq (a : α) : ⊤ ⊓ a = a := inf_of_le_right le_top - --- Porting note: Not simp because simp can prove it -lemma inf_top_eq (a : α) : a ⊓ ⊤ = a := inf_of_le_left le_top - -@[simp] -theorem inf_eq_top_iff : a ⊓ b = ⊤ ↔ a = ⊤ ∧ b = ⊤ := - @sup_eq_bot_iff αᵒᵈ _ _ _ _ - -end SemilatticeInfTop - -section SemilatticeInfBot - -variable [SemilatticeInf α] [OrderBot α] - --- Porting note: Not simp because simp can prove it -lemma bot_inf_eq (a : α) : ⊥ ⊓ a = ⊥ := inf_of_le_left bot_le - --- Porting note: Not simp because simp can prove it -lemma inf_bot_eq (a : α) : a ⊓ ⊥ = ⊥ := inf_of_le_right bot_le - -end SemilatticeInfBot - /-! ### Bounded order -/ @@ -441,101 +350,8 @@ instance BoundedOrder.instSubsingleton : Subsingleton (BoundedOrder α) where end PartialOrder -section Logic - -/-! -#### In this section we prove some properties about monotone and antitone operations on `Prop` --/ - - -section Preorder - -variable [Preorder α] - -theorem monotone_and {p q : α → Prop} (m_p : Monotone p) (m_q : Monotone q) : - Monotone fun x => p x ∧ q x := - fun _ _ h => And.imp (m_p h) (m_q h) - --- Note: by finish [monotone] doesn't work -theorem monotone_or {p q : α → Prop} (m_p : Monotone p) (m_q : Monotone q) : - Monotone fun x => p x ∨ q x := - fun _ _ h => Or.imp (m_p h) (m_q h) - -theorem monotone_le {x : α} : Monotone (x ≤ ·) := fun _ _ h' h => h.trans h' - -theorem monotone_lt {x : α} : Monotone (x < ·) := fun _ _ h' h => h.trans_le h' - -theorem antitone_le {x : α} : Antitone (· ≤ x) := fun _ _ h' h => h'.trans h - -theorem antitone_lt {x : α} : Antitone (· < x) := fun _ _ h' h => h'.trans_lt h - -theorem Monotone.forall {P : β → α → Prop} (hP : ∀ x, Monotone (P x)) : - Monotone fun y => ∀ x, P x y := - fun _ _ hy h x => hP x hy <| h x - -theorem Antitone.forall {P : β → α → Prop} (hP : ∀ x, Antitone (P x)) : - Antitone fun y => ∀ x, P x y := - fun _ _ hy h x => hP x hy (h x) - -theorem Monotone.ball {P : β → α → Prop} {s : Set β} (hP : ∀ x ∈ s, Monotone (P x)) : - Monotone fun y => ∀ x ∈ s, P x y := fun _ _ hy h x hx => hP x hx hy (h x hx) - -theorem Antitone.ball {P : β → α → Prop} {s : Set β} (hP : ∀ x ∈ s, Antitone (P x)) : - Antitone fun y => ∀ x ∈ s, P x y := fun _ _ hy h x hx => hP x hx hy (h x hx) - -theorem Monotone.exists {P : β → α → Prop} (hP : ∀ x, Monotone (P x)) : - Monotone fun y => ∃ x, P x y := - fun _ _ hy ⟨x, hx⟩ ↦ ⟨x, hP x hy hx⟩ - -theorem Antitone.exists {P : β → α → Prop} (hP : ∀ x, Antitone (P x)) : - Antitone fun y => ∃ x, P x y := - fun _ _ hy ⟨x, hx⟩ ↦ ⟨x, hP x hy hx⟩ - -theorem forall_ge_iff {P : α → Prop} {x₀ : α} (hP : Monotone P) : - (∀ x ≥ x₀, P x) ↔ P x₀ := - ⟨fun H ↦ H x₀ le_rfl, fun H _ hx ↦ hP hx H⟩ - -theorem forall_le_iff {P : α → Prop} {x₀ : α} (hP : Antitone P) : - (∀ x ≤ x₀, P x) ↔ P x₀ := - ⟨fun H ↦ H x₀ le_rfl, fun H _ hx ↦ hP hx H⟩ - -end Preorder - -section SemilatticeSup - -variable [SemilatticeSup α] - -theorem exists_ge_and_iff_exists {P : α → Prop} {x₀ : α} (hP : Monotone P) : - (∃ x, x₀ ≤ x ∧ P x) ↔ ∃ x, P x := - ⟨fun h => h.imp fun _ h => h.2, fun ⟨x, hx⟩ => ⟨x ⊔ x₀, le_sup_right, hP le_sup_left hx⟩⟩ - -lemma exists_and_iff_of_monotone {P Q : α → Prop} (hP : Monotone P) (hQ : Monotone Q) : - ((∃ x, P x) ∧ ∃ x, Q x) ↔ (∃ x, P x ∧ Q x) := - ⟨fun ⟨⟨x, hPx⟩, ⟨y, hQx⟩⟩ ↦ ⟨x ⊔ y, ⟨hP le_sup_left hPx, hQ le_sup_right hQx⟩⟩, - fun ⟨x, hPx, hQx⟩ ↦ ⟨⟨x, hPx⟩, ⟨x, hQx⟩⟩⟩ - -end SemilatticeSup - -section SemilatticeInf - -variable [SemilatticeInf α] - -theorem exists_le_and_iff_exists {P : α → Prop} {x₀ : α} (hP : Antitone P) : - (∃ x, x ≤ x₀ ∧ P x) ↔ ∃ x, P x := - exists_ge_and_iff_exists <| hP.dual_left - -lemma exists_and_iff_of_antitone {P Q : α → Prop} (hP : Antitone P) (hQ : Antitone Q) : - ((∃ x, P x) ∧ ∃ x, Q x) ↔ (∃ x, P x ∧ Q x) := - ⟨fun ⟨⟨x, hPx⟩, ⟨y, hQx⟩⟩ ↦ ⟨x ⊓ y, ⟨hP inf_le_left hPx, hQ inf_le_right hQx⟩⟩, - fun ⟨x, hPx, hQx⟩ ↦ ⟨⟨x, hPx⟩, ⟨x, hQx⟩⟩⟩ - -end SemilatticeInf - -end Logic - /-! ### Function lattices -/ - namespace Pi variable {ι : Type*} {α' : ι → Type*} @@ -743,43 +559,6 @@ instance [LE α] [BoundedOrder α] : BoundedOrder (ULift.{v} α) where end ULift -section LinearOrder - -variable [LinearOrder α] - --- `simp` can prove these, so they shouldn't be simp-lemmas. -theorem min_bot_left [OrderBot α] (a : α) : min ⊥ a = ⊥ := bot_inf_eq _ - -theorem max_top_left [OrderTop α] (a : α) : max ⊤ a = ⊤ := top_sup_eq _ - -theorem min_top_left [OrderTop α] (a : α) : min ⊤ a = a := top_inf_eq _ - -theorem max_bot_left [OrderBot α] (a : α) : max ⊥ a = a := bot_sup_eq _ - -theorem min_top_right [OrderTop α] (a : α) : min a ⊤ = a := inf_top_eq _ - -theorem max_bot_right [OrderBot α] (a : α) : max a ⊥ = a := sup_bot_eq _ - -theorem min_bot_right [OrderBot α] (a : α) : min a ⊥ = ⊥ := inf_bot_eq _ - -theorem max_top_right [OrderTop α] (a : α) : max a ⊤ = ⊤ := sup_top_eq _ - -theorem max_eq_bot [OrderBot α] {a b : α} : max a b = ⊥ ↔ a = ⊥ ∧ b = ⊥ := - sup_eq_bot_iff - -theorem min_eq_top [OrderTop α] {a b : α} : min a b = ⊤ ↔ a = ⊤ ∧ b = ⊤ := - inf_eq_top_iff - -@[simp] -theorem min_eq_bot [OrderBot α] {a b : α} : min a b = ⊥ ↔ a = ⊥ ∨ b = ⊥ := by - simp_rw [← le_bot_iff, inf_le_iff] - -@[simp] -theorem max_eq_top [OrderTop α] {a b : α} : max a b = ⊤ ↔ a = ⊤ ∨ b = ⊤ := - @min_eq_bot αᵒᵈ _ _ a b - -end LinearOrder - section Nontrivial variable [PartialOrder α] [BoundedOrder α] [Nontrivial α] diff --git a/Mathlib/Order/BoundedOrder/Lattice.lean b/Mathlib/Order/BoundedOrder/Lattice.lean new file mode 100644 index 0000000000000..42c1cf7c0ac57 --- /dev/null +++ b/Mathlib/Order/BoundedOrder/Lattice.lean @@ -0,0 +1,165 @@ +/- +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 +-/ +import Mathlib.Order.BoundedOrder.Basic +import Mathlib.Order.Lattice + +/-! +# bounded lattices + +This file defines top and bottom elements (greatest and least elements) of a type, the bounded +variants of different kinds of lattices, sets up the typeclass hierarchy between them and provides +instances for `Prop` and `fun`. + +## Common lattices + +* Distributive lattices with a bottom element. Notated by `[DistribLattice α] [OrderBot α]` + It captures the properties of `Disjoint` that are common to `GeneralizedBooleanAlgebra` and + `DistribLattice` when `OrderBot`. +* Bounded and distributive lattice. Notated by `[DistribLattice α] [BoundedOrder α]`. + Typical examples include `Prop` and `Det α`. +-/ + +open Function OrderDual + +universe u v + +variable {α : Type u} {β : Type v} + +/-! ### Top, bottom element -/ + +section SemilatticeSupTop + +variable [SemilatticeSup α] [OrderTop α] + +-- Porting note: Not simp because simp can prove it +theorem top_sup_eq (a : α) : ⊤ ⊔ a = ⊤ := + sup_of_le_left le_top + +-- Porting note: Not simp because simp can prove it +theorem sup_top_eq (a : α) : a ⊔ ⊤ = ⊤ := + sup_of_le_right le_top + +end SemilatticeSupTop + +section SemilatticeSupBot + +variable [SemilatticeSup α] [OrderBot α] {a b : α} + +-- Porting note: Not simp because simp can prove it +theorem bot_sup_eq (a : α) : ⊥ ⊔ a = a := + sup_of_le_right bot_le + +-- Porting note: Not simp because simp can prove it +theorem sup_bot_eq (a : α) : a ⊔ ⊥ = a := + sup_of_le_left bot_le + +@[simp] +theorem sup_eq_bot_iff : a ⊔ b = ⊥ ↔ a = ⊥ ∧ b = ⊥ := by rw [eq_bot_iff, sup_le_iff]; simp + +end SemilatticeSupBot + +section SemilatticeInfTop + +variable [SemilatticeInf α] [OrderTop α] {a b : α} + +-- Porting note: Not simp because simp can prove it +lemma top_inf_eq (a : α) : ⊤ ⊓ a = a := inf_of_le_right le_top + +-- Porting note: Not simp because simp can prove it +lemma inf_top_eq (a : α) : a ⊓ ⊤ = a := inf_of_le_left le_top + +@[simp] +theorem inf_eq_top_iff : a ⊓ b = ⊤ ↔ a = ⊤ ∧ b = ⊤ := + @sup_eq_bot_iff αᵒᵈ _ _ _ _ + +end SemilatticeInfTop + +section SemilatticeInfBot + +variable [SemilatticeInf α] [OrderBot α] + +-- Porting note: Not simp because simp can prove it +lemma bot_inf_eq (a : α) : ⊥ ⊓ a = ⊥ := inf_of_le_left bot_le + +-- Porting note: Not simp because simp can prove it +lemma inf_bot_eq (a : α) : a ⊓ ⊥ = ⊥ := inf_of_le_right bot_le + +end SemilatticeInfBot + +section Logic + +/-! +#### In this section we prove some properties about monotone and antitone operations on `Prop` +-/ + +section SemilatticeSup + +variable [SemilatticeSup α] + +theorem exists_ge_and_iff_exists {P : α → Prop} {x₀ : α} (hP : Monotone P) : + (∃ x, x₀ ≤ x ∧ P x) ↔ ∃ x, P x := + ⟨fun h => h.imp fun _ h => h.2, fun ⟨x, hx⟩ => ⟨x ⊔ x₀, le_sup_right, hP le_sup_left hx⟩⟩ + +lemma exists_and_iff_of_monotone {P Q : α → Prop} (hP : Monotone P) (hQ : Monotone Q) : + ((∃ x, P x) ∧ ∃ x, Q x) ↔ (∃ x, P x ∧ Q x) := + ⟨fun ⟨⟨x, hPx⟩, ⟨y, hQx⟩⟩ ↦ ⟨x ⊔ y, ⟨hP le_sup_left hPx, hQ le_sup_right hQx⟩⟩, + fun ⟨x, hPx, hQx⟩ ↦ ⟨⟨x, hPx⟩, ⟨x, hQx⟩⟩⟩ + +end SemilatticeSup + +section SemilatticeInf + +variable [SemilatticeInf α] + +theorem exists_le_and_iff_exists {P : α → Prop} {x₀ : α} (hP : Antitone P) : + (∃ x, x ≤ x₀ ∧ P x) ↔ ∃ x, P x := + exists_ge_and_iff_exists <| hP.dual_left + +lemma exists_and_iff_of_antitone {P Q : α → Prop} (hP : Antitone P) (hQ : Antitone Q) : + ((∃ x, P x) ∧ ∃ x, Q x) ↔ (∃ x, P x ∧ Q x) := + ⟨fun ⟨⟨x, hPx⟩, ⟨y, hQx⟩⟩ ↦ ⟨x ⊓ y, ⟨hP inf_le_left hPx, hQ inf_le_right hQx⟩⟩, + fun ⟨x, hPx, hQx⟩ ↦ ⟨⟨x, hPx⟩, ⟨x, hQx⟩⟩⟩ + +end SemilatticeInf + +end Logic + +section LinearOrder + +variable [LinearOrder α] + +-- `simp` can prove these, so they shouldn't be simp-lemmas. +theorem min_bot_left [OrderBot α] (a : α) : min ⊥ a = ⊥ := bot_inf_eq _ + +theorem max_top_left [OrderTop α] (a : α) : max ⊤ a = ⊤ := top_sup_eq _ + +theorem min_top_left [OrderTop α] (a : α) : min ⊤ a = a := top_inf_eq _ + +theorem max_bot_left [OrderBot α] (a : α) : max ⊥ a = a := bot_sup_eq _ + +theorem min_top_right [OrderTop α] (a : α) : min a ⊤ = a := inf_top_eq _ + +theorem max_bot_right [OrderBot α] (a : α) : max a ⊥ = a := sup_bot_eq _ + +theorem min_bot_right [OrderBot α] (a : α) : min a ⊥ = ⊥ := inf_bot_eq _ + +theorem max_top_right [OrderTop α] (a : α) : max a ⊤ = ⊤ := sup_top_eq _ + +theorem max_eq_bot [OrderBot α] {a b : α} : max a b = ⊥ ↔ a = ⊥ ∧ b = ⊥ := + sup_eq_bot_iff + +theorem min_eq_top [OrderTop α] {a b : α} : min a b = ⊤ ↔ a = ⊤ ∧ b = ⊤ := + inf_eq_top_iff + +@[simp] +theorem min_eq_bot [OrderBot α] {a b : α} : min a b = ⊥ ↔ a = ⊥ ∨ b = ⊥ := by + simp_rw [← le_bot_iff, inf_le_iff] + +@[simp] +theorem max_eq_top [OrderTop α] {a b : α} : max a b = ⊤ ↔ a = ⊤ ∨ b = ⊤ := + @min_eq_bot αᵒᵈ _ _ a b + +end LinearOrder diff --git a/Mathlib/Order/BoundedOrder/Monotone.lean b/Mathlib/Order/BoundedOrder/Monotone.lean new file mode 100644 index 0000000000000..6b9409de870d5 --- /dev/null +++ b/Mathlib/Order/BoundedOrder/Monotone.lean @@ -0,0 +1,124 @@ +/- +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 +-/ +import Mathlib.Order.BoundedOrder.Basic +import Mathlib.Order.Monotone.Basic + +/-! +# Monotone functions on bounded orders + +-/ + +assert_not_exists SemilatticeSup + +open Function OrderDual + +universe u v + +variable {α : Type u} {β : Type v} + +/-! ### Top, bottom element -/ + +section OrderTop + +variable [PartialOrder α] [OrderTop α] [Preorder β] {f : α → β} {a b : α} + +theorem StrictMono.apply_eq_top_iff (hf : StrictMono f) : f a = f ⊤ ↔ a = ⊤ := + ⟨fun h => not_lt_top_iff.1 fun ha => (hf ha).ne h, congr_arg _⟩ + +theorem StrictAnti.apply_eq_top_iff (hf : StrictAnti f) : f a = f ⊤ ↔ a = ⊤ := + ⟨fun h => not_lt_top_iff.1 fun ha => (hf ha).ne' h, congr_arg _⟩ + +end OrderTop + +theorem StrictMono.maximal_preimage_top [LinearOrder α] [Preorder β] [OrderTop β] {f : α → β} + (H : StrictMono f) {a} (h_top : f a = ⊤) (x : α) : x ≤ a := + H.maximal_of_maximal_image + (fun p => by + rw [h_top] + exact le_top) + x + +section OrderBot + +variable [PartialOrder α] [OrderBot α] [Preorder β] {f : α → β} {a b : α} + +theorem StrictMono.apply_eq_bot_iff (hf : StrictMono f) : f a = f ⊥ ↔ a = ⊥ := + hf.dual.apply_eq_top_iff + +theorem StrictAnti.apply_eq_bot_iff (hf : StrictAnti f) : f a = f ⊥ ↔ a = ⊥ := + hf.dual.apply_eq_top_iff + +end OrderBot + +theorem StrictMono.minimal_preimage_bot [LinearOrder α] [PartialOrder β] [OrderBot β] {f : α → β} + (H : StrictMono f) {a} (h_bot : f a = ⊥) (x : α) : a ≤ x := + H.minimal_of_minimal_image + (fun p => by + rw [h_bot] + exact bot_le) + x + +section Logic + +/-! +#### In this section we prove some properties about monotone and antitone operations on `Prop` +-/ + + +section Preorder + +variable [Preorder α] + +theorem monotone_and {p q : α → Prop} (m_p : Monotone p) (m_q : Monotone q) : + Monotone fun x => p x ∧ q x := + fun _ _ h => And.imp (m_p h) (m_q h) + +-- Note: by finish [monotone] doesn't work +theorem monotone_or {p q : α → Prop} (m_p : Monotone p) (m_q : Monotone q) : + Monotone fun x => p x ∨ q x := + fun _ _ h => Or.imp (m_p h) (m_q h) + +theorem monotone_le {x : α} : Monotone (x ≤ ·) := fun _ _ h' h => h.trans h' + +theorem monotone_lt {x : α} : Monotone (x < ·) := fun _ _ h' h => h.trans_le h' + +theorem antitone_le {x : α} : Antitone (· ≤ x) := fun _ _ h' h => h'.trans h + +theorem antitone_lt {x : α} : Antitone (· < x) := fun _ _ h' h => h'.trans_lt h + +theorem Monotone.forall {P : β → α → Prop} (hP : ∀ x, Monotone (P x)) : + Monotone fun y => ∀ x, P x y := + fun _ _ hy h x => hP x hy <| h x + +theorem Antitone.forall {P : β → α → Prop} (hP : ∀ x, Antitone (P x)) : + Antitone fun y => ∀ x, P x y := + fun _ _ hy h x => hP x hy (h x) + +theorem Monotone.ball {P : β → α → Prop} {s : Set β} (hP : ∀ x ∈ s, Monotone (P x)) : + Monotone fun y => ∀ x ∈ s, P x y := fun _ _ hy h x hx => hP x hx hy (h x hx) + +theorem Antitone.ball {P : β → α → Prop} {s : Set β} (hP : ∀ x ∈ s, Antitone (P x)) : + Antitone fun y => ∀ x ∈ s, P x y := fun _ _ hy h x hx => hP x hx hy (h x hx) + +theorem Monotone.exists {P : β → α → Prop} (hP : ∀ x, Monotone (P x)) : + Monotone fun y => ∃ x, P x y := + fun _ _ hy ⟨x, hx⟩ ↦ ⟨x, hP x hy hx⟩ + +theorem Antitone.exists {P : β → α → Prop} (hP : ∀ x, Antitone (P x)) : + Antitone fun y => ∃ x, P x y := + fun _ _ hy ⟨x, hx⟩ ↦ ⟨x, hP x hy hx⟩ + +theorem forall_ge_iff {P : α → Prop} {x₀ : α} (hP : Monotone P) : + (∀ x ≥ x₀, P x) ↔ P x₀ := + ⟨fun H ↦ H x₀ le_rfl, fun H _ hx ↦ hP hx H⟩ + +theorem forall_le_iff {P : α → Prop} {x₀ : α} (hP : Antitone P) : + (∀ x ≤ x₀, P x) ↔ P x₀ := + ⟨fun H ↦ H x₀ le_rfl, fun H _ hx ↦ hP hx H⟩ + +end Preorder + +end Logic diff --git a/Mathlib/Order/Bounds/Basic.lean b/Mathlib/Order/Bounds/Basic.lean index 99e10ab318f7a..ad2107194dde3 100644 --- a/Mathlib/Order/Bounds/Basic.lean +++ b/Mathlib/Order/Bounds/Basic.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl, Yury Kudryashov -/ import Mathlib.Order.Bounds.Defs import Mathlib.Order.Directed +import Mathlib.Order.BoundedOrder.Monotone import Mathlib.Order.Interval.Set.Basic /-! diff --git a/Mathlib/Order/Disjoint.lean b/Mathlib/Order/Disjoint.lean index 59bcd0815347d..f291e96d7e476 100644 --- a/Mathlib/Order/Disjoint.lean +++ b/Mathlib/Order/Disjoint.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ import Aesop -import Mathlib.Order.BoundedOrder +import Mathlib.Order.BoundedOrder.Lattice /-! # Disjointness and complements diff --git a/Mathlib/Order/Hom/Bounded.lean b/Mathlib/Order/Hom/Bounded.lean index 6466c90ddaf43..feb386498fb37 100644 --- a/Mathlib/Order/Hom/Bounded.lean +++ b/Mathlib/Order/Hom/Bounded.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.Order.BoundedOrder +import Mathlib.Order.BoundedOrder.Lattice /-! # Bounded order homomorphisms diff --git a/Mathlib/Order/Nat.lean b/Mathlib/Order/Nat.lean index 2f527f28b8a53..857c598eeebea 100644 --- a/Mathlib/Order/Nat.lean +++ b/Mathlib/Order/Nat.lean @@ -3,7 +3,7 @@ Copyright (c) 2014 Floris van Doorn (c) 2016 Microsoft Corporation. All rights r Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ -import Mathlib.Order.BoundedOrder +import Mathlib.Order.BoundedOrder.Basic /-! # The natural numbers form a linear order diff --git a/Mathlib/Order/SuccPred/Limit.lean b/Mathlib/Order/SuccPred/Limit.lean index fc51204e9d875..8929a9eebe09a 100644 --- a/Mathlib/Order/SuccPred/Limit.lean +++ b/Mathlib/Order/SuccPred/Limit.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Violeta Hernández Palacios -/ import Mathlib.Order.SuccPred.Archimedean -import Mathlib.Order.BoundedOrder +import Mathlib.Order.BoundedOrder.Lattice /-! # Successor and predecessor limits diff --git a/Mathlib/Order/WithBot.lean b/Mathlib/Order/WithBot.lean index 5be2bf90e1157..da8552cf675af 100644 --- a/Mathlib/Order/WithBot.lean +++ b/Mathlib/Order/WithBot.lean @@ -4,12 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ import Mathlib.Logic.Nontrivial.Basic -import Mathlib.Order.BoundedOrder import Mathlib.Order.TypeTags import Mathlib.Data.Option.NAry import Mathlib.Tactic.Contrapose import Mathlib.Tactic.Lift import Mathlib.Data.Option.Basic +import Mathlib.Order.Lattice +import Mathlib.Order.BoundedOrder.Basic /-! # `WithBot`, `WithTop` diff --git a/scripts/noshake.json b/scripts/noshake.json index 4a06c0c7680d4..b3c884f2347c8 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -337,6 +337,7 @@ "Mathlib.Order.Filter.ListTraverse": ["Mathlib.Control.Traversable.Instances"], "Mathlib.Order.Defs": ["Batteries.Tactic.Trans"], + "Mathlib.Order.BoundedOrder.Basic": ["Mathlib.Tactic.Finiteness.Attr"], "Mathlib.Order.BoundedOrder": ["Mathlib.Tactic.Finiteness.Attr"], "Mathlib.Order.Basic": ["Batteries.Tactic.Classical", "Mathlib.Tactic.GCongr.Core"], From a2a544ab0cc7d3e1d2a07cda2e9dedef8dbe0749 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 26 Nov 2024 07:18:19 +0000 Subject: [PATCH 311/829] chore: move OrderDual.instSup/Inf up to Order.Basic (#19489) --- Mathlib/Order/Basic.lean | 6 ++++++ Mathlib/Order/Lattice.lean | 6 ------ 2 files changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Order/Basic.lean b/Mathlib/Order/Basic.lean index 945ab1df613af..7812d48426cce 100644 --- a/Mathlib/Order/Basic.lean +++ b/Mathlib/Order/Basic.lean @@ -691,6 +691,12 @@ instance (α : Type*) [LE α] : LE αᵒᵈ := instance (α : Type*) [LT α] : LT αᵒᵈ := ⟨fun x y : α ↦ y < x⟩ +instance instSup (α : Type*) [Min α] : Max αᵒᵈ := + ⟨((· ⊓ ·) : α → α → α)⟩ + +instance instInf (α : Type*) [Max α] : Min αᵒᵈ := + ⟨((· ⊔ ·) : α → α → α)⟩ + instance instPreorder (α : Type*) [Preorder α] : Preorder αᵒᵈ where le_refl := fun _ ↦ le_refl _ le_trans := fun _ _ _ hab hbc ↦ hbc.trans hab diff --git a/Mathlib/Order/Lattice.lean b/Mathlib/Order/Lattice.lean index f141ab6888510..acfa2bc47f35f 100644 --- a/Mathlib/Order/Lattice.lean +++ b/Mathlib/Order/Lattice.lean @@ -95,12 +95,6 @@ def SemilatticeSup.mk' {α : Type*} [Max α] (sup_comm : ∀ a b : α, a ⊔ b = le_sup_right a b := by dsimp; rw [sup_comm, sup_assoc, sup_idem] sup_le a b c hac hbc := by dsimp; rwa [sup_assoc, hbc] -instance OrderDual.instSup (α : Type*) [Min α] : Max αᵒᵈ := - ⟨((· ⊓ ·) : α → α → α)⟩ - -instance OrderDual.instInf (α : Type*) [Max α] : Min αᵒᵈ := - ⟨((· ⊔ ·) : α → α → α)⟩ - section SemilatticeSup variable [SemilatticeSup α] {a b c d : α} From 1ebeab334af93de65153b3fbb8708777bab5149b Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 26 Nov 2024 08:42:19 +0000 Subject: [PATCH 312/829] chore: move le_refl's simp attribute to definition (#19490) --- Mathlib/Order/Basic.lean | 2 -- Mathlib/Order/Defs.lean | 2 +- 2 files changed, 1 insertion(+), 3 deletions(-) diff --git a/Mathlib/Order/Basic.lean b/Mathlib/Order/Basic.lean index 7812d48426cce..84bcd9f95e9fe 100644 --- a/Mathlib/Order/Basic.lean +++ b/Mathlib/Order/Basic.lean @@ -98,8 +98,6 @@ theorem Ne.lt_of_le' : b ≠ a → a ≤ b → a < b := end PartialOrder -attribute [simp] le_refl - attribute [ext] LE alias LE.le.trans := le_trans diff --git a/Mathlib/Order/Defs.lean b/Mathlib/Order/Defs.lean index 8c0f11d2df7c6..31c1c7508d026 100644 --- a/Mathlib/Order/Defs.lean +++ b/Mathlib/Order/Defs.lean @@ -264,7 +264,7 @@ class Preorder (α : Type*) extends LE α, LT α where variable [Preorder α] {a b c : α} /-- The relation `≤` on a preorder is reflexive. -/ -@[refl] lemma le_refl : ∀ a : α, a ≤ a := Preorder.le_refl +@[refl, simp] lemma le_refl : ∀ a : α, a ≤ a := Preorder.le_refl /-- A version of `le_refl` where the argument is implicit -/ lemma le_rfl : a ≤ a := le_refl a From 83238e24b9113e84b6e4a66801a7dc8c4e9c8320 Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 26 Nov 2024 08:42:21 +0000 Subject: [PATCH 313/829] fix: `bors x` comments and zulip reactions (#19491) Restores the bot reacting to `bors d`, `bors merge` and `bors r+`. It only searches in `PR reviews`, but at least it should no longer fail CI. I plan to investigate making the bot search all of Zulip in a later PR. --- scripts/zulip_emoji_merge_delegate.py | 96 +++++++++++++-------------- 1 file changed, 46 insertions(+), 50 deletions(-) diff --git a/scripts/zulip_emoji_merge_delegate.py b/scripts/zulip_emoji_merge_delegate.py index b279fbc6e1e11..e572aeb90e8d3 100644 --- a/scripts/zulip_emoji_merge_delegate.py +++ b/scripts/zulip_emoji_merge_delegate.py @@ -25,74 +25,70 @@ # Fetch the last 200 messages response = client.get_messages({ - "operator": "search", - "operand": f"https://github\.com/leanprover-community/mathlib4/pull/{PR_NUMBER}", + "anchor": "newest", + "num_before": 200, + "num_after": 0, + "narrow": [{"operator": "channel", "operand": "PR reviews"}], }) messages = response['messages'] +pr_pattern = re.compile(f'https://github.com/leanprover-community/mathlib4/pull/{PR_NUMBER}') + +print(f"Searching for: '{pr_pattern}'") + for message in messages: content = message['content'] - print(f"matched: '{message}'") # Check for emoji reactions reactions = message['reactions'] has_peace_sign = any(reaction['emoji_name'] == 'peace_sign' for reaction in reactions) has_bors = any(reaction['emoji_name'] == 'bors' for reaction in reactions) has_merge = any(reaction['emoji_name'] == 'merge' for reaction in reactions) + match = pr_pattern.search(content) + if match: + print(f"matched: '{message}'") - pr_url = f"https://api.github.com/repos/leanprover-community/mathlib4/pulls/{PR_NUMBER}" - - print('Removing peace_sign') - result = client.remove_reaction({ - "message_id": message['id'], - "emoji_name": "peace_sign" - }) - print(f"result: '{result}'") - print('Removing bors') - result = client.remove_reaction({ - "message_id": message['id'], - "emoji_name": "bors" - }) - print(f"result: '{result}'") - - print('Removing merge') - result = client.remove_reaction({ - "message_id": message['id'], - "emoji_name": "merge" - }) - print(f"result: '{result}'") - - if 'delegated' == LABEL: - print('adding delegated') - - client.add_reaction({ - "message_id": message['id'], - "emoji_name": "peace_sign" - }) - elif 'ready-to-merge' == LABEL: - print('adding ready-to-merge') + # removing previous emoji reactions + print("Removing previous reactions, if present.") if has_peace_sign: - client.remove_reaction({ + print('Removing peace_sign') + result = client.remove_reaction({ "message_id": message['id'], "emoji_name": "peace_sign" }) - client.add_reaction({ - "message_id": message['id'], - "emoji_name": "bors" - }) - elif LABEL.startswith("[Merged by Bors]"): - print('adding [Merged by Bors]') - if has_peace_sign: - client.remove_reaction({ + print(f"result: '{result}'") + if has_bors: + print('Removing bors') + result = client.remove_reaction({ "message_id": message['id'], - "emoji_name": "peace_sign" + "emoji_name": "bors" }) - if has_bors: - client.remove_reaction({ + print(f"result: '{result}'") + if has_merge: + print('Removing merge') + result = client.remove_reaction({ + "message_id": message['id'], + "emoji_name": "merge" + }) + print(f"result: '{result}'") + + # applying appropriate emoji reaction + print("Applying reactions, as appropriate.") + if 'ready-to-merge' == LABEL: + print('adding ready-to-merge') + client.add_reaction({ "message_id": message['id'], "emoji_name": "bors" }) - client.add_reaction({ - "message_id": message['id'], - "emoji_name": "merge" - }) + elif 'delegated' == LABEL: + print('adding delegated') + client.add_reaction({ + "message_id": message['id'], + "emoji_name": "peace_sign" + }) + elif LABEL.startswith("[Merged by Bors]"): + print('adding [Merged by Bors]') + client.add_reaction({ + "message_id": message['id'], + "emoji_name": "merge" + }) From 6c84e8f7efc21583aca881d7c9a8ce9ea25f740c Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 26 Nov 2024 10:41:42 +0000 Subject: [PATCH 314/829] chore: split Order.Defs (#19498) --- Mathlib.lean | 4 +- Mathlib/Algebra/Group/Commute/Defs.lean | 1 - Mathlib/Algebra/Group/Semiconj/Defs.lean | 2 +- Mathlib/Algebra/NeZero.lean | 2 +- Mathlib/Data/Bool/Basic.lean | 1 + Mathlib/Data/Char.lean | 2 +- Mathlib/Data/Int/Order/Basic.lean | 2 +- Mathlib/Data/Ordering/Lemmas.lean | 2 +- Mathlib/Deprecated/AlgebraClasses.lean | 3 +- Mathlib/Deprecated/Order.lean | 3 +- Mathlib/Logic/Basic.lean | 7 +- Mathlib/Logic/Function/Basic.lean | 2 + Mathlib/Logic/Nontrivial/Basic.lean | 2 +- Mathlib/Order/Basic.lean | 2 +- Mathlib/Order/Defs.lean | 708 ----------------------- Mathlib/Order/Defs/LinearOrder.lean | 311 ++++++++++ Mathlib/Order/Defs/PartialOrder.lean | 158 +++++ Mathlib/Order/Defs/Unbundled.lean | 278 +++++++++ Mathlib/Order/Interval/Set/Defs.lean | 2 +- Mathlib/Tactic/GCongr/Core.lean | 6 +- Mathlib/Tactic/PushNeg.lean | 2 +- MathlibTest/push_neg.lean | 2 +- scripts/noshake.json | 5 +- 23 files changed, 775 insertions(+), 732 deletions(-) delete mode 100644 Mathlib/Order/Defs.lean create mode 100644 Mathlib/Order/Defs/LinearOrder.lean create mode 100644 Mathlib/Order/Defs/PartialOrder.lean create mode 100644 Mathlib/Order/Defs/Unbundled.lean diff --git a/Mathlib.lean b/Mathlib.lean index fddf683280feb..c81570dc36eb0 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3935,7 +3935,9 @@ import Mathlib.Order.ConditionallyCompleteLattice.Indexed import Mathlib.Order.Copy import Mathlib.Order.CountableDenseLinearOrder import Mathlib.Order.Cover -import Mathlib.Order.Defs +import Mathlib.Order.Defs.LinearOrder +import Mathlib.Order.Defs.PartialOrder +import Mathlib.Order.Defs.Unbundled import Mathlib.Order.Directed import Mathlib.Order.DirectedInverseSystem import Mathlib.Order.Disjoint diff --git a/Mathlib/Algebra/Group/Commute/Defs.lean b/Mathlib/Algebra/Group/Commute/Defs.lean index 07568fa4a02eb..4d77aadfd19ab 100644 --- a/Mathlib/Algebra/Group/Commute/Defs.lean +++ b/Mathlib/Algebra/Group/Commute/Defs.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Neil Strickland, Yury Kudryashov -/ import Mathlib.Algebra.Group.Semiconj.Defs -import Mathlib.Order.Defs /-! # Commuting pairs of elements in monoids diff --git a/Mathlib/Algebra/Group/Semiconj/Defs.lean b/Mathlib/Algebra/Group/Semiconj/Defs.lean index cfeabeb0d12db..dc45b2afef400 100644 --- a/Mathlib/Algebra/Group/Semiconj/Defs.lean +++ b/Mathlib/Algebra/Group/Semiconj/Defs.lean @@ -6,7 +6,7 @@ Authors: Yury Kudryashov -- Some proofs and docs came from mathlib3 `src/algebra/commute.lean` (c) Neil Strickland import Mathlib.Algebra.Group.Defs -import Mathlib.Order.Defs +import Mathlib.Order.Defs.Unbundled /-! # Semiconjugate elements of a semigroup diff --git a/Mathlib/Algebra/NeZero.lean b/Mathlib/Algebra/NeZero.lean index 3fdf3e370cb38..35ba9754734ce 100644 --- a/Mathlib/Algebra/NeZero.lean +++ b/Mathlib/Algebra/NeZero.lean @@ -5,7 +5,7 @@ Authors: Eric Rodriguez -/ import Mathlib.Logic.Basic import Mathlib.Algebra.Group.ZeroOne -import Mathlib.Order.Defs +import Mathlib.Order.Defs.PartialOrder /-! # `NeZero` typeclass diff --git a/Mathlib/Data/Bool/Basic.lean b/Mathlib/Data/Bool/Basic.lean index 11a7e6829bea0..1b5e3ff38f835 100644 --- a/Mathlib/Data/Bool/Basic.lean +++ b/Mathlib/Data/Bool/Basic.lean @@ -5,6 +5,7 @@ Authors: Leonardo de Moura, Jeremy Avigad -/ import Mathlib.Logic.Basic import Mathlib.Logic.Function.Defs +import Mathlib.Order.Defs.LinearOrder /-! # Booleans diff --git a/Mathlib/Data/Char.lean b/Mathlib/Data/Char.lean index 69187fc4b9eb6..ebc7f9e688601 100644 --- a/Mathlib/Data/Char.lean +++ b/Mathlib/Data/Char.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Mathlib.Data.Nat.Defs -import Mathlib.Order.Defs +import Mathlib.Order.Defs.LinearOrder /-! # More `Char` instances diff --git a/Mathlib/Data/Int/Order/Basic.lean b/Mathlib/Data/Int/Order/Basic.lean index 4fcb9a0a8f8f8..04758c04bf6f5 100644 --- a/Mathlib/Data/Int/Order/Basic.lean +++ b/Mathlib/Data/Int/Order/Basic.lean @@ -6,7 +6,7 @@ Authors: Jeremy Avigad import Mathlib.Data.Int.Notation import Mathlib.Data.Nat.Notation -import Mathlib.Order.Defs +import Mathlib.Order.Defs.LinearOrder /-! # The order relation on the integers diff --git a/Mathlib/Data/Ordering/Lemmas.lean b/Mathlib/Data/Ordering/Lemmas.lean index 7b9714aa85ed5..42da3c30c0fb9 100644 --- a/Mathlib/Data/Ordering/Lemmas.lean +++ b/Mathlib/Data/Ordering/Lemmas.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura -/ import Mathlib.Data.Ordering.Basic -import Mathlib.Order.Defs +import Mathlib.Order.Defs.Unbundled /-! # Some `Ordering` lemmas diff --git a/Mathlib/Deprecated/AlgebraClasses.lean b/Mathlib/Deprecated/AlgebraClasses.lean index bb6f948871913..e9e879b6c6967 100644 --- a/Mathlib/Deprecated/AlgebraClasses.lean +++ b/Mathlib/Deprecated/AlgebraClasses.lean @@ -3,7 +3,8 @@ Copyright (c) 2017 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura -/ -import Mathlib.Order.Defs +import Mathlib.Order.Defs.Unbundled +import Mathlib.Order.Defs.LinearOrder /-! # Note about deprecated files diff --git a/Mathlib/Deprecated/Order.lean b/Mathlib/Deprecated/Order.lean index 8e0ef3256036e..9f79797c7e46f 100644 --- a/Mathlib/Deprecated/Order.lean +++ b/Mathlib/Deprecated/Order.lean @@ -3,7 +3,8 @@ Copyright (c) 2014 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Floris van Doorn -/ -import Mathlib.Order.Defs +import Mathlib.Order.Defs.Unbundled +import Mathlib.Order.Defs.LinearOrder /-! # Deprecated instances about order structures. diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index 89c3a8fe8a2f1..13c8fd29ba91a 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -6,11 +6,10 @@ Authors: Jeremy Avigad, Leonardo de Moura import Mathlib.Tactic.Attr.Register import Mathlib.Tactic.Basic import Batteries.Logic +import Batteries.Tactic.Trans import Batteries.Util.LibraryNote -import Batteries.Tactic.Lint.Basic import Mathlib.Data.Nat.Notation import Mathlib.Data.Int.Notation -import Mathlib.Order.Defs /-! # Basic logic properties @@ -133,10 +132,6 @@ section Propositional /-! ### Declarations about `implies` -/ -instance : IsRefl Prop Iff := ⟨Iff.refl⟩ - -instance : IsTrans Prop Iff := ⟨fun _ _ _ ↦ Iff.trans⟩ - alias Iff.imp := imp_congr -- This is a duplicate of `Classical.imp_iff_right_iff`. Deprecate? diff --git a/Mathlib/Logic/Function/Basic.lean b/Mathlib/Logic/Function/Basic.lean index f9b8452f77c2b..5f794b63872c6 100644 --- a/Mathlib/Logic/Function/Basic.lean +++ b/Mathlib/Logic/Function/Basic.lean @@ -8,6 +8,8 @@ import Mathlib.Logic.Basic import Mathlib.Logic.ExistsUnique import Mathlib.Logic.Nonempty import Batteries.Tactic.Init +import Mathlib.Order.Defs.PartialOrder +import Mathlib.Order.Defs.Unbundled /-! # Miscellaneous function constructions and lemmas diff --git a/Mathlib/Logic/Nontrivial/Basic.lean b/Mathlib/Logic/Nontrivial/Basic.lean index ee2fcbb4f7759..987edfff6dd39 100644 --- a/Mathlib/Logic/Nontrivial/Basic.lean +++ b/Mathlib/Logic/Nontrivial/Basic.lean @@ -7,7 +7,7 @@ import Mathlib.Data.Prod.Basic import Mathlib.Logic.Function.Basic import Mathlib.Logic.Nontrivial.Defs import Mathlib.Logic.Unique -import Mathlib.Order.Defs +import Mathlib.Order.Defs.LinearOrder import Mathlib.Tactic.Attr.Register /-! diff --git a/Mathlib/Order/Basic.lean b/Mathlib/Order/Basic.lean index 84bcd9f95e9fe..caa54da5b5024 100644 --- a/Mathlib/Order/Basic.lean +++ b/Mathlib/Order/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Mario Carneiro -/ import Mathlib.Data.Subtype -import Mathlib.Order.Defs +import Mathlib.Order.Defs.LinearOrder import Mathlib.Order.Notation import Mathlib.Tactic.GCongr.Core import Mathlib.Tactic.Spread diff --git a/Mathlib/Order/Defs.lean b/Mathlib/Order/Defs.lean deleted file mode 100644 index 31c1c7508d026..0000000000000 --- a/Mathlib/Order/Defs.lean +++ /dev/null @@ -1,708 +0,0 @@ -/- -Copyright (c) 2016 Microsoft Corporation. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Leonardo de Moura --/ -import Batteries.Classes.Order -import Batteries.Tactic.Trans -import Mathlib.Data.Ordering.Basic -import Mathlib.Data.Set.Defs -import Mathlib.Tactic.ExtendDoc -import Mathlib.Tactic.Lemma -import Mathlib.Tactic.SplitIfs -import Mathlib.Tactic.TypeStar - -/-! -# Orders - -Defines classes for preorders, partial orders, and linear orders -and proves some basic lemmas about them. --/ - -/-! ### Unbundled classes -/ - -/-- An empty relation does not relate any elements. -/ -@[nolint unusedArguments] def EmptyRelation {α : Sort*} := fun _ _ : α ↦ False - -/-- `IsIrrefl X r` means the binary relation `r` on `X` is irreflexive (that is, `r x x` never -holds). -/ -class IsIrrefl (α : Sort*) (r : α → α → Prop) : Prop where - irrefl : ∀ a, ¬r a a - -/-- `IsRefl X r` means the binary relation `r` on `X` is reflexive. -/ -class IsRefl (α : Sort*) (r : α → α → Prop) : Prop where - refl : ∀ a, r a a - -/-- `IsSymm X r` means the binary relation `r` on `X` is symmetric. -/ -class IsSymm (α : Sort*) (r : α → α → Prop) : Prop where - symm : ∀ a b, r a b → r b a - -/-- `IsAsymm X r` means that the binary relation `r` on `X` is asymmetric, that is, -`r a b → ¬ r b a`. -/ -class IsAsymm (α : Sort*) (r : α → α → Prop) : Prop where - asymm : ∀ a b, r a b → ¬r b a - -/-- `IsAntisymm X r` means the binary relation `r` on `X` is antisymmetric. -/ -class IsAntisymm (α : Sort*) (r : α → α → Prop) : Prop where - antisymm : ∀ a b, r a b → r b a → a = b - -instance (priority := 100) IsAsymm.toIsAntisymm {α : Sort*} (r : α → α → Prop) [IsAsymm α r] : - IsAntisymm α r where - antisymm _ _ hx hy := (IsAsymm.asymm _ _ hx hy).elim - -/-- `IsTrans X r` means the binary relation `r` on `X` is transitive. -/ -class IsTrans (α : Sort*) (r : α → α → Prop) : Prop where - trans : ∀ a b c, r a b → r b c → r a c - -instance {α : Sort*} {r : α → α → Prop} [IsTrans α r] : Trans r r r := - ⟨IsTrans.trans _ _ _⟩ - -instance (priority := 100) {α : Sort*} {r : α → α → Prop} [Trans r r r] : IsTrans α r := - ⟨fun _ _ _ => Trans.trans⟩ - -/-- `IsTotal X r` means that the binary relation `r` on `X` is total, that is, that for any -`x y : X` we have `r x y` or `r y x`. -/ -class IsTotal (α : Sort*) (r : α → α → Prop) : Prop where - total : ∀ a b, r a b ∨ r b a - -/-- `IsPreorder X r` means that the binary relation `r` on `X` is a pre-order, that is, reflexive -and transitive. -/ -class IsPreorder (α : Sort*) (r : α → α → Prop) extends IsRefl α r, IsTrans α r : Prop - -/-- `IsPartialOrder X r` means that the binary relation `r` on `X` is a partial order, that is, -`IsPreorder X r` and `IsAntisymm X r`. -/ -class IsPartialOrder (α : Sort*) (r : α → α → Prop) extends IsPreorder α r, IsAntisymm α r : Prop - -/-- `IsLinearOrder X r` means that the binary relation `r` on `X` is a linear order, that is, -`IsPartialOrder X r` and `IsTotal X r`. -/ -class IsLinearOrder (α : Sort*) (r : α → α → Prop) extends IsPartialOrder α r, IsTotal α r : Prop - -/-- `IsEquiv X r` means that the binary relation `r` on `X` is an equivalence relation, that -is, `IsPreorder X r` and `IsSymm X r`. -/ -class IsEquiv (α : Sort*) (r : α → α → Prop) extends IsPreorder α r, IsSymm α r : Prop - -/-- `IsStrictOrder X r` means that the binary relation `r` on `X` is a strict order, that is, -`IsIrrefl X r` and `IsTrans X r`. -/ -class IsStrictOrder (α : Sort*) (r : α → α → Prop) extends IsIrrefl α r, IsTrans α r : Prop - -/-- `IsStrictWeakOrder X lt` means that the binary relation `lt` on `X` is a strict weak order, -that is, `IsStrictOrder X lt` and `¬lt a b ∧ ¬lt b a → ¬lt b c ∧ ¬lt c b → ¬lt a c ∧ ¬lt c a`. -/ -class IsStrictWeakOrder (α : Sort*) (lt : α → α → Prop) extends IsStrictOrder α lt : Prop where - incomp_trans : ∀ a b c, ¬lt a b ∧ ¬lt b a → ¬lt b c ∧ ¬lt c b → ¬lt a c ∧ ¬lt c a - -/-- `IsTrichotomous X lt` means that the binary relation `lt` on `X` is trichotomous, that is, -either `lt a b` or `a = b` or `lt b a` for any `a` and `b`. -/ -class IsTrichotomous (α : Sort*) (lt : α → α → Prop) : Prop where - trichotomous : ∀ a b, lt a b ∨ a = b ∨ lt b a - -/-- `IsStrictTotalOrder X lt` means that the binary relation `lt` on `X` is a strict total order, -that is, `IsTrichotomous X lt` and `IsStrictOrder X lt`. -/ -class IsStrictTotalOrder (α : Sort*) (lt : α → α → Prop) extends IsTrichotomous α lt, - IsStrictOrder α lt : Prop - -/-- Equality is an equivalence relation. -/ -instance eq_isEquiv (α : Sort*) : IsEquiv α (· = ·) where - symm := @Eq.symm _ - trans := @Eq.trans _ - refl := Eq.refl - -section - -variable {α : Sort*} {r : α → α → Prop} {a b c : α} - -/-- Local notation for an arbitrary binary relation `r`. -/ -local infixl:50 " ≺ " => r - -lemma irrefl [IsIrrefl α r] (a : α) : ¬a ≺ a := IsIrrefl.irrefl a -lemma refl [IsRefl α r] (a : α) : a ≺ a := IsRefl.refl a -lemma trans [IsTrans α r] : a ≺ b → b ≺ c → a ≺ c := IsTrans.trans _ _ _ -lemma symm [IsSymm α r] : a ≺ b → b ≺ a := IsSymm.symm _ _ -lemma antisymm [IsAntisymm α r] : a ≺ b → b ≺ a → a = b := IsAntisymm.antisymm _ _ -lemma asymm [IsAsymm α r] : a ≺ b → ¬b ≺ a := IsAsymm.asymm _ _ - -lemma trichotomous [IsTrichotomous α r] : ∀ a b : α, a ≺ b ∨ a = b ∨ b ≺ a := - IsTrichotomous.trichotomous - -instance (priority := 90) isAsymm_of_isTrans_of_isIrrefl [IsTrans α r] [IsIrrefl α r] : - IsAsymm α r := - ⟨fun a _b h₁ h₂ => absurd (_root_.trans h₁ h₂) (irrefl a)⟩ - -instance IsIrrefl.decide [DecidableRel r] [IsIrrefl α r] : - IsIrrefl α (fun a b => decide (r a b) = true) where - irrefl := fun a => by simpa using irrefl a - -instance IsRefl.decide [DecidableRel r] [IsRefl α r] : - IsRefl α (fun a b => decide (r a b) = true) where - refl := fun a => by simpa using refl a - -instance IsTrans.decide [DecidableRel r] [IsTrans α r] : - IsTrans α (fun a b => decide (r a b) = true) where - trans := fun a b c => by simpa using trans a b c - -instance IsSymm.decide [DecidableRel r] [IsSymm α r] : - IsSymm α (fun a b => decide (r a b) = true) where - symm := fun a b => by simpa using symm a b - -instance IsAntisymm.decide [DecidableRel r] [IsAntisymm α r] : - IsAntisymm α (fun a b => decide (r a b) = true) where - antisymm := fun a b h₁ h₂ => antisymm _ _ (by simpa using h₁) (by simpa using h₂) - -instance IsAsymm.decide [DecidableRel r] [IsAsymm α r] : - IsAsymm α (fun a b => decide (r a b) = true) where - asymm := fun a b => by simpa using asymm a b - -instance IsTotal.decide [DecidableRel r] [IsTotal α r] : - IsTotal α (fun a b => decide (r a b) = true) where - total := fun a b => by simpa using total a b - -instance IsTrichotomous.decide [DecidableRel r] [IsTrichotomous α r] : - IsTrichotomous α (fun a b => decide (r a b) = true) where - trichotomous := fun a b => by simpa using trichotomous a b - -variable (r) - -@[elab_without_expected_type] lemma irrefl_of [IsIrrefl α r] (a : α) : ¬a ≺ a := irrefl a -@[elab_without_expected_type] lemma refl_of [IsRefl α r] (a : α) : a ≺ a := refl a -@[elab_without_expected_type] lemma trans_of [IsTrans α r] : a ≺ b → b ≺ c → a ≺ c := _root_.trans -@[elab_without_expected_type] lemma symm_of [IsSymm α r] : a ≺ b → b ≺ a := symm -@[elab_without_expected_type] lemma asymm_of [IsAsymm α r] : a ≺ b → ¬b ≺ a := asymm - -@[elab_without_expected_type] -lemma total_of [IsTotal α r] (a b : α) : a ≺ b ∨ b ≺ a := IsTotal.total _ _ - -@[elab_without_expected_type] -lemma trichotomous_of [IsTrichotomous α r] : ∀ a b : α, a ≺ b ∨ a = b ∨ b ≺ a := trichotomous - -section - -/-- `IsRefl` as a definition, suitable for use in proofs. -/ -def Reflexive := ∀ x, x ≺ x - -/-- `IsSymm` as a definition, suitable for use in proofs. -/ -def Symmetric := ∀ ⦃x y⦄, x ≺ y → y ≺ x - -/-- `IsTrans` as a definition, suitable for use in proofs. -/ -def Transitive := ∀ ⦃x y z⦄, x ≺ y → y ≺ z → x ≺ z - -/-- `IsIrrefl` as a definition, suitable for use in proofs. -/ -def Irreflexive := ∀ x, ¬x ≺ x - -/-- `IsAntisymm` as a definition, suitable for use in proofs. -/ -def AntiSymmetric := ∀ ⦃x y⦄, x ≺ y → y ≺ x → x = y - -/-- `IsTotal` as a definition, suitable for use in proofs. -/ -def Total := ∀ x y, x ≺ y ∨ y ≺ x - -@[deprecated Equivalence.refl (since := "2024-09-13")] -theorem Equivalence.reflexive (h : Equivalence r) : Reflexive r := h.refl - -@[deprecated Equivalence.symm (since := "2024-09-13")] -theorem Equivalence.symmetric (h : Equivalence r) : Symmetric r := - fun _ _ ↦ h.symm - -@[deprecated Equivalence.trans (since := "2024-09-13")] -theorem Equivalence.transitive (h : Equivalence r) : Transitive r := - fun _ _ _ ↦ h.trans - -variable {β : Sort*} (r : β → β → Prop) (f : α → β) - -@[deprecated "No deprecation message was provided." (since := "2024-09-13")] -theorem InvImage.trans (h : Transitive r) : Transitive (InvImage r f) := - fun (a₁ a₂ a₃ : α) (h₁ : InvImage r f a₁ a₂) (h₂ : InvImage r f a₂ a₃) ↦ h h₁ h₂ - -@[deprecated "No deprecation message was provided." (since := "2024-09-13")] -theorem InvImage.irreflexive (h : Irreflexive r) : Irreflexive (InvImage r f) := - fun (a : α) (h₁ : InvImage r f a a) ↦ h (f a) h₁ - -end - -end - -/-! ### Minimal and maximal -/ - -section LE - -variable {α : Type*} [LE α] {P : α → Prop} {x y : α} - -/-- `Minimal P x` means that `x` is a minimal element satisfying `P`. -/ -def Minimal (P : α → Prop) (x : α) : Prop := P x ∧ ∀ ⦃y⦄, P y → y ≤ x → x ≤ y - -/-- `Maximal P x` means that `x` is a maximal element satisfying `P`. -/ -def Maximal (P : α → Prop) (x : α) : Prop := P x ∧ ∀ ⦃y⦄, P y → x ≤ y → y ≤ x - -lemma Minimal.prop (h : Minimal P x) : P x := - h.1 - -lemma Maximal.prop (h : Maximal P x) : P x := - h.1 - -lemma Minimal.le_of_le (h : Minimal P x) (hy : P y) (hle : y ≤ x) : x ≤ y := - h.2 hy hle - -lemma Maximal.le_of_ge (h : Maximal P x) (hy : P y) (hge : x ≤ y) : y ≤ x := - h.2 hy hge - -end LE - -/-! ### Bundled classes -/ - -variable {α : Type*} - -section Preorder - -/-! -### Definition of `Preorder` and lemmas about types with a `Preorder` --/ - -/-- A preorder is a reflexive, transitive relation `≤` with `a < b` defined in the obvious way. -/ -class Preorder (α : Type*) extends LE α, LT α where - le_refl : ∀ a : α, a ≤ a - le_trans : ∀ a b c : α, a ≤ b → b ≤ c → a ≤ c - lt := fun a b => a ≤ b ∧ ¬b ≤ a - lt_iff_le_not_le : ∀ a b : α, a < b ↔ a ≤ b ∧ ¬b ≤ a := by intros; rfl - -variable [Preorder α] {a b c : α} - -/-- The relation `≤` on a preorder is reflexive. -/ -@[refl, simp] lemma le_refl : ∀ a : α, a ≤ a := Preorder.le_refl - -/-- A version of `le_refl` where the argument is implicit -/ -lemma le_rfl : a ≤ a := le_refl a - -/-- The relation `≤` on a preorder is transitive. -/ -@[trans] lemma le_trans : a ≤ b → b ≤ c → a ≤ c := Preorder.le_trans _ _ _ - -lemma lt_iff_le_not_le : a < b ↔ a ≤ b ∧ ¬b ≤ a := Preorder.lt_iff_le_not_le _ _ - -lemma lt_of_le_not_le (hab : a ≤ b) (hba : ¬ b ≤ a) : a < b := lt_iff_le_not_le.2 ⟨hab, hba⟩ - -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -theorem le_not_le_of_lt : ∀ {a b : α}, a < b → a ≤ b ∧ ¬b ≤ a - | _a, _b, hab => lt_iff_le_not_le.mp hab - -lemma le_of_eq (hab : a = b) : a ≤ b := by rw [hab] -lemma le_of_lt (hab : a < b) : a ≤ b := (lt_iff_le_not_le.1 hab).1 -lemma not_le_of_lt (hab : a < b) : ¬ b ≤ a := (lt_iff_le_not_le.1 hab).2 -lemma not_le_of_gt (hab : a > b) : ¬a ≤ b := not_le_of_lt hab -lemma not_lt_of_le (hab : a ≤ b) : ¬ b < a := imp_not_comm.1 not_le_of_lt hab -lemma not_lt_of_ge (hab : a ≥ b) : ¬a < b := not_lt_of_le hab - -alias LT.lt.not_le := not_le_of_lt -alias LE.le.not_lt := not_lt_of_le - -@[trans] lemma ge_trans : a ≥ b → b ≥ c → a ≥ c := fun h₁ h₂ => le_trans h₂ h₁ - -lemma lt_irrefl (a : α) : ¬a < a := fun h ↦ not_le_of_lt h le_rfl -lemma gt_irrefl (a : α) : ¬a > a := lt_irrefl _ - -@[trans] lemma lt_of_lt_of_le (hab : a < b) (hbc : b ≤ c) : a < c := - lt_of_le_not_le (le_trans (le_of_lt hab) hbc) fun hca ↦ not_le_of_lt hab (le_trans hbc hca) - -@[trans] lemma lt_of_le_of_lt (hab : a ≤ b) (hbc : b < c) : a < c := - lt_of_le_not_le (le_trans hab (le_of_lt hbc)) fun hca ↦ not_le_of_lt hbc (le_trans hca hab) - -@[trans] lemma gt_of_gt_of_ge (h₁ : a > b) (h₂ : b ≥ c) : a > c := lt_of_le_of_lt h₂ h₁ -@[trans] lemma gt_of_ge_of_gt (h₁ : a ≥ b) (h₂ : b > c) : a > c := lt_of_lt_of_le h₂ h₁ - -@[trans] lemma lt_trans (hab : a < b) (hbc : b < c) : a < c := lt_of_lt_of_le hab (le_of_lt hbc) -@[trans] lemma gt_trans : a > b → b > c → a > c := fun h₁ h₂ => lt_trans h₂ h₁ - -lemma ne_of_lt (h : a < b) : a ≠ b := fun he => absurd h (he ▸ lt_irrefl a) -lemma ne_of_gt (h : b < a) : a ≠ b := fun he => absurd h (he ▸ lt_irrefl a) -lemma lt_asymm (h : a < b) : ¬b < a := fun h1 : b < a => lt_irrefl a (lt_trans h h1) - -alias not_lt_of_gt := lt_asymm -alias not_lt_of_lt := lt_asymm - -lemma le_of_lt_or_eq (h : a < b ∨ a = b) : a ≤ b := h.elim le_of_lt le_of_eq -lemma le_of_eq_or_lt (h : a = b ∨ a < b) : a ≤ b := h.elim le_of_eq le_of_lt - -instance (priority := 900) : @Trans α α α LE.le LE.le LE.le := ⟨le_trans⟩ -instance (priority := 900) : @Trans α α α LT.lt LT.lt LT.lt := ⟨lt_trans⟩ -instance (priority := 900) : @Trans α α α LT.lt LE.le LT.lt := ⟨lt_of_lt_of_le⟩ -instance (priority := 900) : @Trans α α α LE.le LT.lt LT.lt := ⟨lt_of_le_of_lt⟩ -instance (priority := 900) : @Trans α α α GE.ge GE.ge GE.ge := ⟨ge_trans⟩ -instance (priority := 900) : @Trans α α α GT.gt GT.gt GT.gt := ⟨gt_trans⟩ -instance (priority := 900) : @Trans α α α GT.gt GE.ge GT.gt := ⟨gt_of_gt_of_ge⟩ -instance (priority := 900) : @Trans α α α GE.ge GT.gt GT.gt := ⟨gt_of_ge_of_gt⟩ - -/-- `<` is decidable if `≤` is. -/ -def decidableLTOfDecidableLE [@DecidableRel α (· ≤ ·)] : @DecidableRel α (· < ·) - | a, b => - if hab : a ≤ b then - if hba : b ≤ a then isFalse fun hab' => not_le_of_gt hab' hba - else isTrue <| lt_of_le_not_le hab hba - else isFalse fun hab' => hab (le_of_lt hab') - -end Preorder - -section PartialOrder - -/-! -### Definition of `PartialOrder` and lemmas about types with a partial order --/ - -/-- A partial order is a reflexive, transitive, antisymmetric relation `≤`. -/ -class PartialOrder (α : Type*) extends Preorder α where - le_antisymm : ∀ a b : α, a ≤ b → b ≤ a → a = b - -variable [PartialOrder α] {a b : α} - -lemma le_antisymm : a ≤ b → b ≤ a → a = b := PartialOrder.le_antisymm _ _ - -alias eq_of_le_of_le := le_antisymm - -lemma le_antisymm_iff : a = b ↔ a ≤ b ∧ b ≤ a := - ⟨fun e => ⟨le_of_eq e, le_of_eq e.symm⟩, fun ⟨h1, h2⟩ => le_antisymm h1 h2⟩ - -lemma lt_of_le_of_ne : a ≤ b → a ≠ b → a < b := fun h₁ h₂ => - lt_of_le_not_le h₁ <| mt (le_antisymm h₁) h₂ - -/-- Equality is decidable if `≤` is. -/ -def decidableEqOfDecidableLE [@DecidableRel α (· ≤ ·)] : DecidableEq α - | a, b => - if hab : a ≤ b then - if hba : b ≤ a then isTrue (le_antisymm hab hba) else isFalse fun heq => hba (heq ▸ le_refl _) - else isFalse fun heq => hab (heq ▸ le_refl _) - -namespace Decidable - -variable [@DecidableRel α (· ≤ ·)] - -lemma lt_or_eq_of_le (hab : a ≤ b) : a < b ∨ a = b := - if hba : b ≤ a then Or.inr (le_antisymm hab hba) else Or.inl (lt_of_le_not_le hab hba) - -lemma eq_or_lt_of_le (hab : a ≤ b) : a = b ∨ a < b := - (lt_or_eq_of_le hab).symm - -lemma le_iff_lt_or_eq : a ≤ b ↔ a < b ∨ a = b := - ⟨lt_or_eq_of_le, le_of_lt_or_eq⟩ - -end Decidable - -attribute [local instance] Classical.propDecidable - -lemma lt_or_eq_of_le : a ≤ b → a < b ∨ a = b := Decidable.lt_or_eq_of_le -lemma le_iff_lt_or_eq : a ≤ b ↔ a < b ∨ a = b := Decidable.le_iff_lt_or_eq - -end PartialOrder - -section LinearOrder - -/-! -### Definition of `LinearOrder` and lemmas about types with a linear order --/ - -/-- Default definition of `max`. -/ -def maxDefault [LE α] [DecidableRel ((· ≤ ·) : α → α → Prop)] (a b : α) := - if a ≤ b then b else a - -/-- Default definition of `min`. -/ -def minDefault [LE α] [DecidableRel ((· ≤ ·) : α → α → Prop)] (a b : α) := - if a ≤ b then a else b - -/-- This attempts to prove that a given instance of `compare` is equal to `compareOfLessAndEq` by -introducing the arguments and trying the following approaches in order: - -1. seeing if `rfl` works -2. seeing if the `compare` at hand is nonetheless essentially `compareOfLessAndEq`, but, because of -implicit arguments, requires us to unfold the defs and split the `if`s in the definition of -`compareOfLessAndEq` -3. seeing if we can split by cases on the arguments, then see if the defs work themselves out - (useful when `compare` is defined via a `match` statement, as it is for `Bool`) -/ -macro "compareOfLessAndEq_rfl" : tactic => - `(tactic| (intros a b; first | rfl | - (simp only [compare, compareOfLessAndEq]; split_ifs <;> rfl) | - (induction a <;> induction b <;> simp +decide only))) - -/-- A linear order is reflexive, transitive, antisymmetric and total relation `≤`. -We assume that every linear ordered type has decidable `(≤)`, `(<)`, and `(=)`. -/ -class LinearOrder (α : Type*) extends PartialOrder α, Min α, Max α, Ord α where - /-- A linear order is total. -/ - le_total (a b : α) : a ≤ b ∨ b ≤ a - /-- In a linearly ordered type, we assume the order relations are all decidable. -/ - decidableLE : DecidableRel (· ≤ · : α → α → Prop) - /-- In a linearly ordered type, we assume the order relations are all decidable. -/ - decidableEq : DecidableEq α := @decidableEqOfDecidableLE _ _ decidableLE - /-- In a linearly ordered type, we assume the order relations are all decidable. -/ - decidableLT : DecidableRel (· < · : α → α → Prop) := - @decidableLTOfDecidableLE _ _ decidableLE - min := fun a b => if a ≤ b then a else b - max := fun a b => if a ≤ b then b else a - /-- The minimum function is equivalent to the one you get from `minOfLe`. -/ - min_def : ∀ a b, min a b = if a ≤ b then a else b := by intros; rfl - /-- The minimum function is equivalent to the one you get from `maxOfLe`. -/ - max_def : ∀ a b, max a b = if a ≤ b then b else a := by intros; rfl - compare a b := compareOfLessAndEq a b - /-- Comparison via `compare` is equal to the canonical comparison given decidable `<` and `=`. -/ - compare_eq_compareOfLessAndEq : ∀ a b, compare a b = compareOfLessAndEq a b := by - compareOfLessAndEq_rfl - -variable [LinearOrder α] {a b c : α} - -attribute [local instance] LinearOrder.decidableLE - -lemma le_total : ∀ a b : α, a ≤ b ∨ b ≤ a := LinearOrder.le_total - -lemma le_of_not_ge : ¬a ≥ b → a ≤ b := (le_total b a).resolve_left -lemma le_of_not_le : ¬a ≤ b → b ≤ a := (le_total a b).resolve_left -lemma lt_of_not_ge (h : ¬a ≥ b) : a < b := lt_of_le_not_le (le_of_not_ge h) h - -lemma lt_trichotomy (a b : α) : a < b ∨ a = b ∨ b < a := - Or.elim (le_total a b) - (fun h : a ≤ b => - Or.elim (Decidable.lt_or_eq_of_le h) (fun h : a < b => Or.inl h) fun h : a = b => - Or.inr (Or.inl h)) - fun h : b ≤ a => - Or.elim (Decidable.lt_or_eq_of_le h) (fun h : b < a => Or.inr (Or.inr h)) fun h : b = a => - Or.inr (Or.inl h.symm) - -lemma le_of_not_lt (h : ¬b < a) : a ≤ b := - match lt_trichotomy a b with - | Or.inl hlt => le_of_lt hlt - | Or.inr (Or.inl HEq) => HEq ▸ le_refl a - | Or.inr (Or.inr hgt) => absurd hgt h - -lemma le_of_not_gt : ¬a > b → a ≤ b := le_of_not_lt - -lemma lt_or_le (a b : α) : a < b ∨ b ≤ a := - if hba : b ≤ a then Or.inr hba else Or.inl <| lt_of_not_ge hba - -lemma le_or_lt (a b : α) : a ≤ b ∨ b < a := (lt_or_le b a).symm -lemma lt_or_ge : ∀ a b : α, a < b ∨ a ≥ b := lt_or_le -lemma le_or_gt : ∀ a b : α, a ≤ b ∨ a > b := le_or_lt - -lemma lt_or_gt_of_ne (h : a ≠ b) : a < b ∨ a > b := by simpa [h] using lt_trichotomy a b - -lemma ne_iff_lt_or_gt : a ≠ b ↔ a < b ∨ a > b := ⟨lt_or_gt_of_ne, (Or.elim · ne_of_lt ne_of_gt)⟩ - -lemma lt_iff_not_ge (x y : α) : x < y ↔ ¬x ≥ y := ⟨not_le_of_gt, lt_of_not_ge⟩ - -@[simp] lemma not_lt : ¬a < b ↔ b ≤ a := ⟨le_of_not_gt, not_lt_of_ge⟩ -@[simp] lemma not_le : ¬a ≤ b ↔ b < a := (lt_iff_not_ge _ _).symm - -instance (priority := 900) (a b : α) : Decidable (a < b) := LinearOrder.decidableLT a b -instance (priority := 900) (a b : α) : Decidable (a ≤ b) := LinearOrder.decidableLE a b -instance (priority := 900) (a b : α) : Decidable (a = b) := LinearOrder.decidableEq a b - -lemma eq_or_lt_of_not_lt (h : ¬a < b) : a = b ∨ b < a := - if h₁ : a = b then Or.inl h₁ else Or.inr (lt_of_not_ge fun hge => h (lt_of_le_of_ne hge h₁)) - -/-- Perform a case-split on the ordering of `x` and `y` in a decidable linear order. -/ -def ltByCases (x y : α) {P : Sort*} (h₁ : x < y → P) (h₂ : x = y → P) (h₃ : y < x → P) : P := - if h : x < y then h₁ h - else if h' : y < x then h₃ h' else h₂ (le_antisymm (le_of_not_gt h') (le_of_not_gt h)) - -namespace Nat - -/-! Deprecated properties of inequality on `Nat` -/ - -@[deprecated "No deprecation message was provided." (since := "2024-08-23")] -protected def ltGeByCases {a b : Nat} {C : Sort*} (h₁ : a < b → C) (h₂ : b ≤ a → C) : C := - Decidable.byCases h₁ fun h => h₂ (Or.elim (Nat.lt_or_ge a b) (fun a => absurd a h) fun a => a) - -set_option linter.deprecated false in -@[deprecated ltByCases (since := "2024-08-23")] -protected def ltByCases {a b : Nat} {C : Sort*} (h₁ : a < b → C) (h₂ : a = b → C) - (h₃ : b < a → C) : C := - Nat.ltGeByCases h₁ fun h₁ => Nat.ltGeByCases h₃ fun h => h₂ (Nat.le_antisymm h h₁) - -end Nat - -theorem le_imp_le_of_lt_imp_lt {α β} [Preorder α] [LinearOrder β] {a b : α} {c d : β} - (H : d < c → b < a) (h : a ≤ b) : c ≤ d := - le_of_not_lt fun h' => not_le_of_gt (H h') h - -lemma min_def (a b : α) : min a b = if a ≤ b then a else b := by rw [LinearOrder.min_def a] -lemma max_def (a b : α) : max a b = if a ≤ b then b else a := by rw [LinearOrder.max_def a] - --- Porting note: no `min_tac` tactic in the following series of lemmas - -lemma min_le_left (a b : α) : min a b ≤ a := by - if h : a ≤ b - then simp [min_def, if_pos h, le_refl] - else simp [min_def, if_neg h]; exact le_of_not_le h - -lemma min_le_right (a b : α) : min a b ≤ b := by - if h : a ≤ b - then simp [min_def, if_pos h]; exact h - else simp [min_def, if_neg h, le_refl] - -lemma le_min (h₁ : c ≤ a) (h₂ : c ≤ b) : c ≤ min a b := by - if h : a ≤ b - then simp [min_def, if_pos h]; exact h₁ - else simp [min_def, if_neg h]; exact h₂ - -lemma le_max_left (a b : α) : a ≤ max a b := by - if h : a ≤ b - then simp [max_def, if_pos h]; exact h - else simp [max_def, if_neg h, le_refl] - -lemma le_max_right (a b : α) : b ≤ max a b := by - if h : a ≤ b - then simp [max_def, if_pos h, le_refl] - else simp [max_def, if_neg h]; exact le_of_not_le h - -lemma max_le (h₁ : a ≤ c) (h₂ : b ≤ c) : max a b ≤ c := by - if h : a ≤ b - then simp [max_def, if_pos h]; exact h₂ - else simp [max_def, if_neg h]; exact h₁ - -lemma eq_min (h₁ : c ≤ a) (h₂ : c ≤ b) (h₃ : ∀ {d}, d ≤ a → d ≤ b → d ≤ c) : c = min a b := - le_antisymm (le_min h₁ h₂) (h₃ (min_le_left a b) (min_le_right a b)) - -lemma min_comm (a b : α) : min a b = min b a := - eq_min (min_le_right a b) (min_le_left a b) fun h₁ h₂ => le_min h₂ h₁ - -lemma min_assoc (a b c : α) : min (min a b) c = min a (min b c) := by - apply eq_min - · apply le_trans (min_le_left ..); apply min_le_left - · apply le_min - · 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 - · 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] - -@[simp] lemma min_self (a : α) : min a a = a := by simp [min_def] - -lemma min_eq_left (h : a ≤ b) : min a b = a := by - apply Eq.symm; apply eq_min (le_refl _) h; intros; assumption - -lemma min_eq_right (h : b ≤ a) : min a b = b := min_comm b a ▸ min_eq_left h - -lemma eq_max (h₁ : a ≤ c) (h₂ : b ≤ c) (h₃ : ∀ {d}, a ≤ d → b ≤ d → c ≤ d) : - c = max a b := - le_antisymm (h₃ (le_max_left a b) (le_max_right a b)) (max_le h₁ h₂) - -lemma max_comm (a b : α) : max a b = max b a := - eq_max (le_max_right a b) (le_max_left a b) fun h₁ h₂ => max_le h₂ h₁ - -lemma max_assoc (a b c : α) : max (max a b) c = max a (max b c) := by - apply eq_max - · apply le_trans (le_max_left a b); apply le_max_left - · apply max_le - · 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₂ - · 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] - -@[simp] lemma max_self (a : α) : max a a = a := by simp [max_def] - -lemma max_eq_left (h : b ≤ a) : max a b = a := by - apply Eq.symm; apply eq_max (le_refl _) h; intros; assumption - -lemma max_eq_right (h : a ≤ b) : max a b = b := max_comm b a ▸ max_eq_left h - -lemma min_eq_left_of_lt (h : a < b) : min a b = a := min_eq_left (le_of_lt h) -lemma min_eq_right_of_lt (h : b < a) : min a b = b := min_eq_right (le_of_lt h) -lemma max_eq_left_of_lt (h : b < a) : max a b = a := max_eq_left (le_of_lt h) -lemma max_eq_right_of_lt (h : a < b) : max a b = b := max_eq_right (le_of_lt h) - -lemma lt_min (h₁ : a < b) (h₂ : a < c) : a < min b c := by - cases le_total b c <;> simp [min_eq_left, min_eq_right, *] - -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 - -lemma compare_lt_iff_lt : compare a b = .lt ↔ a < b := by - rw [LinearOrder.compare_eq_compareOfLessAndEq, compareOfLessAndEq] - split_ifs <;> simp only [*, lt_irrefl, reduceCtorEq] - -lemma compare_gt_iff_gt : compare a b = .gt ↔ a > b := by - rw [LinearOrder.compare_eq_compareOfLessAndEq, compareOfLessAndEq] - split_ifs <;> simp only [*, lt_irrefl, not_lt_of_gt, reduceCtorEq] - case _ h₁ h₂ => - have h : b < a := lt_trichotomy a b |>.resolve_left h₁ |>.resolve_left h₂ - rwa [true_iff] - -lemma compare_eq_iff_eq : compare a b = .eq ↔ a = b := by - rw [LinearOrder.compare_eq_compareOfLessAndEq, compareOfLessAndEq] - split_ifs <;> try simp only [reduceCtorEq] - case _ h => rw [false_iff]; exact ne_iff_lt_or_gt.2 <| .inl h - case _ _ h => rwa [true_iff] - case _ _ h => rwa [false_iff] - -lemma compare_le_iff_le : compare a b ≠ .gt ↔ a ≤ b := by - cases h : compare a b <;> simp - · exact le_of_lt <| compare_lt_iff_lt.1 h - · exact le_of_eq <| compare_eq_iff_eq.1 h - · exact compare_gt_iff_gt.1 h - -lemma compare_ge_iff_ge : compare a b ≠ .lt ↔ a ≥ b := by - cases h : compare a b <;> simp - · exact compare_lt_iff_lt.1 h - · exact le_of_eq <| (·.symm) <| compare_eq_iff_eq.1 h - · exact le_of_lt <| compare_gt_iff_gt.1 h - -lemma compare_iff (a b : α) {o : Ordering} : compare a b = o ↔ o.Compares a b := by - cases o <;> simp only [Ordering.Compares] - · exact compare_lt_iff_lt - · exact compare_eq_iff_eq - · exact compare_gt_iff_gt - -theorem cmp_eq_compare (a b : α) : cmp a b = compare a b := by - refine ((compare_iff ..).2 ?_).symm - unfold cmp cmpUsing; split_ifs with h1 h2 - · exact h1 - · exact h2 - · exact le_antisymm (not_lt.1 h2) (not_lt.1 h1) - -theorem cmp_eq_compareOfLessAndEq (a b : α) : cmp a b = compareOfLessAndEq a b := - (cmp_eq_compare ..).trans (LinearOrder.compare_eq_compareOfLessAndEq ..) - -instance : Batteries.LawfulCmp (compare (α := α)) where - symm a b := by - cases h : compare a b <;> - simp only [Ordering.swap] <;> symm - · exact compare_gt_iff_gt.2 <| compare_lt_iff_lt.1 h - · exact compare_eq_iff_eq.2 <| compare_eq_iff_eq.1 h |>.symm - · exact compare_lt_iff_lt.2 <| compare_gt_iff_gt.1 h - le_trans := fun h₁ h₂ ↦ - compare_le_iff_le.2 <| le_trans (compare_le_iff_le.1 h₁) (compare_le_iff_le.1 h₂) - cmp_iff_beq := by simp [compare_eq_iff_eq] - cmp_iff_lt := by simp [compare_lt_iff_lt] - cmp_iff_le := by simp [compare_le_iff_le] - -end Ord - -end LinearOrder - -/-! ### Upper and lower sets -/ - -/-- An upper set in an order `α` is a set such that any element greater than one of its members is -also a member. Also called up-set, upward-closed set. -/ -def IsUpperSet {α : Type*} [LE α] (s : Set α) : Prop := - ∀ ⦃a b : α⦄, a ≤ b → a ∈ s → b ∈ s - -/-- A lower set in an order `α` is a set such that any element less than one of its members is also -a member. Also called down-set, downward-closed set. -/ -def IsLowerSet {α : Type*} [LE α] (s : Set α) : Prop := - ∀ ⦃a b : α⦄, b ≤ a → a ∈ s → b ∈ s - -@[inherit_doc IsUpperSet] -structure UpperSet (α : Type*) [LE α] where - /-- The carrier of an `UpperSet`. -/ - carrier : Set α - /-- The carrier of an `UpperSet` is an upper set. -/ - upper' : IsUpperSet carrier - -extend_docs UpperSet before "The type of upper sets of an order." - -@[inherit_doc IsLowerSet] -structure LowerSet (α : Type*) [LE α] where - /-- The carrier of a `LowerSet`. -/ - carrier : Set α - /-- The carrier of a `LowerSet` is a lower set. -/ - lower' : IsLowerSet carrier - -extend_docs LowerSet before "The type of lower sets of an order." diff --git a/Mathlib/Order/Defs/LinearOrder.lean b/Mathlib/Order/Defs/LinearOrder.lean new file mode 100644 index 0000000000000..ec281a950f1d6 --- /dev/null +++ b/Mathlib/Order/Defs/LinearOrder.lean @@ -0,0 +1,311 @@ +/- +Copyright (c) 2016 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura +-/ +import Batteries.Classes.Order +import Batteries.Tactic.Trans +import Mathlib.Data.Ordering.Basic +import Mathlib.Tactic.ExtendDoc +import Mathlib.Tactic.Lemma +import Mathlib.Tactic.SplitIfs +import Mathlib.Tactic.TypeStar +import Mathlib.Order.Defs.PartialOrder + +/-! +# Orders + +Defines classes for linear orders and proves some basic lemmas about them. +-/ + +variable {α : Type*} + +section LinearOrder + +/-! +### Definition of `LinearOrder` and lemmas about types with a linear order +-/ + +/-- Default definition of `max`. -/ +def maxDefault [LE α] [DecidableRel ((· ≤ ·) : α → α → Prop)] (a b : α) := + if a ≤ b then b else a + +/-- Default definition of `min`. -/ +def minDefault [LE α] [DecidableRel ((· ≤ ·) : α → α → Prop)] (a b : α) := + if a ≤ b then a else b + +/-- This attempts to prove that a given instance of `compare` is equal to `compareOfLessAndEq` by +introducing the arguments and trying the following approaches in order: + +1. seeing if `rfl` works +2. seeing if the `compare` at hand is nonetheless essentially `compareOfLessAndEq`, but, because of +implicit arguments, requires us to unfold the defs and split the `if`s in the definition of +`compareOfLessAndEq` +3. seeing if we can split by cases on the arguments, then see if the defs work themselves out + (useful when `compare` is defined via a `match` statement, as it is for `Bool`) -/ +macro "compareOfLessAndEq_rfl" : tactic => + `(tactic| (intros a b; first | rfl | + (simp only [compare, compareOfLessAndEq]; split_ifs <;> rfl) | + (induction a <;> induction b <;> simp +decide only))) + +/-- A linear order is reflexive, transitive, antisymmetric and total relation `≤`. +We assume that every linear ordered type has decidable `(≤)`, `(<)`, and `(=)`. -/ +class LinearOrder (α : Type*) extends PartialOrder α, Min α, Max α, Ord α where + /-- A linear order is total. -/ + le_total (a b : α) : a ≤ b ∨ b ≤ a + /-- In a linearly ordered type, we assume the order relations are all decidable. -/ + decidableLE : DecidableRel (· ≤ · : α → α → Prop) + /-- In a linearly ordered type, we assume the order relations are all decidable. -/ + decidableEq : DecidableEq α := @decidableEqOfDecidableLE _ _ decidableLE + /-- In a linearly ordered type, we assume the order relations are all decidable. -/ + decidableLT : DecidableRel (· < · : α → α → Prop) := + @decidableLTOfDecidableLE _ _ decidableLE + min := fun a b => if a ≤ b then a else b + max := fun a b => if a ≤ b then b else a + /-- The minimum function is equivalent to the one you get from `minOfLe`. -/ + min_def : ∀ a b, min a b = if a ≤ b then a else b := by intros; rfl + /-- The minimum function is equivalent to the one you get from `maxOfLe`. -/ + max_def : ∀ a b, max a b = if a ≤ b then b else a := by intros; rfl + compare a b := compareOfLessAndEq a b + /-- Comparison via `compare` is equal to the canonical comparison given decidable `<` and `=`. -/ + compare_eq_compareOfLessAndEq : ∀ a b, compare a b = compareOfLessAndEq a b := by + compareOfLessAndEq_rfl + +variable [LinearOrder α] {a b c : α} + +attribute [local instance] LinearOrder.decidableLE + +lemma le_total : ∀ a b : α, a ≤ b ∨ b ≤ a := LinearOrder.le_total + +lemma le_of_not_ge : ¬a ≥ b → a ≤ b := (le_total b a).resolve_left +lemma le_of_not_le : ¬a ≤ b → b ≤ a := (le_total a b).resolve_left +lemma lt_of_not_ge (h : ¬a ≥ b) : a < b := lt_of_le_not_le (le_of_not_ge h) h + +lemma lt_trichotomy (a b : α) : a < b ∨ a = b ∨ b < a := + Or.elim (le_total a b) + (fun h : a ≤ b => + Or.elim (Decidable.lt_or_eq_of_le h) (fun h : a < b => Or.inl h) fun h : a = b => + Or.inr (Or.inl h)) + fun h : b ≤ a => + Or.elim (Decidable.lt_or_eq_of_le h) (fun h : b < a => Or.inr (Or.inr h)) fun h : b = a => + Or.inr (Or.inl h.symm) + +lemma le_of_not_lt (h : ¬b < a) : a ≤ b := + match lt_trichotomy a b with + | Or.inl hlt => le_of_lt hlt + | Or.inr (Or.inl HEq) => HEq ▸ le_refl a + | Or.inr (Or.inr hgt) => absurd hgt h + +lemma le_of_not_gt : ¬a > b → a ≤ b := le_of_not_lt + +lemma lt_or_le (a b : α) : a < b ∨ b ≤ a := + if hba : b ≤ a then Or.inr hba else Or.inl <| lt_of_not_ge hba + +lemma le_or_lt (a b : α) : a ≤ b ∨ b < a := (lt_or_le b a).symm +lemma lt_or_ge : ∀ a b : α, a < b ∨ a ≥ b := lt_or_le +lemma le_or_gt : ∀ a b : α, a ≤ b ∨ a > b := le_or_lt + +lemma lt_or_gt_of_ne (h : a ≠ b) : a < b ∨ a > b := by simpa [h] using lt_trichotomy a b + +lemma ne_iff_lt_or_gt : a ≠ b ↔ a < b ∨ a > b := ⟨lt_or_gt_of_ne, (Or.elim · ne_of_lt ne_of_gt)⟩ + +lemma lt_iff_not_ge (x y : α) : x < y ↔ ¬x ≥ y := ⟨not_le_of_gt, lt_of_not_ge⟩ + +@[simp] lemma not_lt : ¬a < b ↔ b ≤ a := ⟨le_of_not_gt, not_lt_of_ge⟩ +@[simp] lemma not_le : ¬a ≤ b ↔ b < a := (lt_iff_not_ge _ _).symm + +instance (priority := 900) (a b : α) : Decidable (a < b) := LinearOrder.decidableLT a b +instance (priority := 900) (a b : α) : Decidable (a ≤ b) := LinearOrder.decidableLE a b +instance (priority := 900) (a b : α) : Decidable (a = b) := LinearOrder.decidableEq a b + +lemma eq_or_lt_of_not_lt (h : ¬a < b) : a = b ∨ b < a := + if h₁ : a = b then Or.inl h₁ else Or.inr (lt_of_not_ge fun hge => h (lt_of_le_of_ne hge h₁)) + +/-- Perform a case-split on the ordering of `x` and `y` in a decidable linear order. -/ +def ltByCases (x y : α) {P : Sort*} (h₁ : x < y → P) (h₂ : x = y → P) (h₃ : y < x → P) : P := + if h : x < y then h₁ h + else if h' : y < x then h₃ h' else h₂ (le_antisymm (le_of_not_gt h') (le_of_not_gt h)) + +namespace Nat + +/-! Deprecated properties of inequality on `Nat` -/ + +@[deprecated "No deprecation message was provided." (since := "2024-08-23")] +protected def ltGeByCases {a b : Nat} {C : Sort*} (h₁ : a < b → C) (h₂ : b ≤ a → C) : C := + Decidable.byCases h₁ fun h => h₂ (Or.elim (Nat.lt_or_ge a b) (fun a => absurd a h) fun a => a) + +set_option linter.deprecated false in +@[deprecated ltByCases (since := "2024-08-23")] +protected def ltByCases {a b : Nat} {C : Sort*} (h₁ : a < b → C) (h₂ : a = b → C) + (h₃ : b < a → C) : C := + Nat.ltGeByCases h₁ fun h₁ => Nat.ltGeByCases h₃ fun h => h₂ (Nat.le_antisymm h h₁) + +end Nat + +theorem le_imp_le_of_lt_imp_lt {α β} [Preorder α] [LinearOrder β] {a b : α} {c d : β} + (H : d < c → b < a) (h : a ≤ b) : c ≤ d := + le_of_not_lt fun h' => not_le_of_gt (H h') h + +lemma min_def (a b : α) : min a b = if a ≤ b then a else b := by rw [LinearOrder.min_def a] +lemma max_def (a b : α) : max a b = if a ≤ b then b else a := by rw [LinearOrder.max_def a] + +-- Porting note: no `min_tac` tactic in the following series of lemmas + +lemma min_le_left (a b : α) : min a b ≤ a := by + if h : a ≤ b + then simp [min_def, if_pos h, le_refl] + else simp [min_def, if_neg h]; exact le_of_not_le h + +lemma min_le_right (a b : α) : min a b ≤ b := by + if h : a ≤ b + then simp [min_def, if_pos h]; exact h + else simp [min_def, if_neg h, le_refl] + +lemma le_min (h₁ : c ≤ a) (h₂ : c ≤ b) : c ≤ min a b := by + if h : a ≤ b + then simp [min_def, if_pos h]; exact h₁ + else simp [min_def, if_neg h]; exact h₂ + +lemma le_max_left (a b : α) : a ≤ max a b := by + if h : a ≤ b + then simp [max_def, if_pos h]; exact h + else simp [max_def, if_neg h, le_refl] + +lemma le_max_right (a b : α) : b ≤ max a b := by + if h : a ≤ b + then simp [max_def, if_pos h, le_refl] + else simp [max_def, if_neg h]; exact le_of_not_le h + +lemma max_le (h₁ : a ≤ c) (h₂ : b ≤ c) : max a b ≤ c := by + if h : a ≤ b + then simp [max_def, if_pos h]; exact h₂ + else simp [max_def, if_neg h]; exact h₁ + +lemma eq_min (h₁ : c ≤ a) (h₂ : c ≤ b) (h₃ : ∀ {d}, d ≤ a → d ≤ b → d ≤ c) : c = min a b := + le_antisymm (le_min h₁ h₂) (h₃ (min_le_left a b) (min_le_right a b)) + +lemma min_comm (a b : α) : min a b = min b a := + eq_min (min_le_right a b) (min_le_left a b) fun h₁ h₂ => le_min h₂ h₁ + +lemma min_assoc (a b c : α) : min (min a b) c = min a (min b c) := by + apply eq_min + · apply le_trans (min_le_left ..); apply min_le_left + · apply le_min + · 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 + · 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] + +@[simp] lemma min_self (a : α) : min a a = a := by simp [min_def] + +lemma min_eq_left (h : a ≤ b) : min a b = a := by + apply Eq.symm; apply eq_min (le_refl _) h; intros; assumption + +lemma min_eq_right (h : b ≤ a) : min a b = b := min_comm b a ▸ min_eq_left h + +lemma eq_max (h₁ : a ≤ c) (h₂ : b ≤ c) (h₃ : ∀ {d}, a ≤ d → b ≤ d → c ≤ d) : + c = max a b := + le_antisymm (h₃ (le_max_left a b) (le_max_right a b)) (max_le h₁ h₂) + +lemma max_comm (a b : α) : max a b = max b a := + eq_max (le_max_right a b) (le_max_left a b) fun h₁ h₂ => max_le h₂ h₁ + +lemma max_assoc (a b c : α) : max (max a b) c = max a (max b c) := by + apply eq_max + · apply le_trans (le_max_left a b); apply le_max_left + · apply max_le + · 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₂ + · 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] + +@[simp] lemma max_self (a : α) : max a a = a := by simp [max_def] + +lemma max_eq_left (h : b ≤ a) : max a b = a := by + apply Eq.symm; apply eq_max (le_refl _) h; intros; assumption + +lemma max_eq_right (h : a ≤ b) : max a b = b := max_comm b a ▸ max_eq_left h + +lemma min_eq_left_of_lt (h : a < b) : min a b = a := min_eq_left (le_of_lt h) +lemma min_eq_right_of_lt (h : b < a) : min a b = b := min_eq_right (le_of_lt h) +lemma max_eq_left_of_lt (h : b < a) : max a b = a := max_eq_left (le_of_lt h) +lemma max_eq_right_of_lt (h : a < b) : max a b = b := max_eq_right (le_of_lt h) + +lemma lt_min (h₁ : a < b) (h₂ : a < c) : a < min b c := by + cases le_total b c <;> simp [min_eq_left, min_eq_right, *] + +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 + +lemma compare_lt_iff_lt : compare a b = .lt ↔ a < b := by + rw [LinearOrder.compare_eq_compareOfLessAndEq, compareOfLessAndEq] + split_ifs <;> simp only [*, lt_irrefl, reduceCtorEq] + +lemma compare_gt_iff_gt : compare a b = .gt ↔ a > b := by + rw [LinearOrder.compare_eq_compareOfLessAndEq, compareOfLessAndEq] + split_ifs <;> simp only [*, lt_irrefl, not_lt_of_gt, reduceCtorEq] + case _ h₁ h₂ => + have h : b < a := lt_trichotomy a b |>.resolve_left h₁ |>.resolve_left h₂ + rwa [true_iff] + +lemma compare_eq_iff_eq : compare a b = .eq ↔ a = b := by + rw [LinearOrder.compare_eq_compareOfLessAndEq, compareOfLessAndEq] + split_ifs <;> try simp only [reduceCtorEq] + case _ h => rw [false_iff]; exact ne_iff_lt_or_gt.2 <| .inl h + case _ _ h => rwa [true_iff] + case _ _ h => rwa [false_iff] + +lemma compare_le_iff_le : compare a b ≠ .gt ↔ a ≤ b := by + cases h : compare a b <;> simp + · exact le_of_lt <| compare_lt_iff_lt.1 h + · exact le_of_eq <| compare_eq_iff_eq.1 h + · exact compare_gt_iff_gt.1 h + +lemma compare_ge_iff_ge : compare a b ≠ .lt ↔ a ≥ b := by + cases h : compare a b <;> simp + · exact compare_lt_iff_lt.1 h + · exact le_of_eq <| (·.symm) <| compare_eq_iff_eq.1 h + · exact le_of_lt <| compare_gt_iff_gt.1 h + +lemma compare_iff (a b : α) {o : Ordering} : compare a b = o ↔ o.Compares a b := by + cases o <;> simp only [Ordering.Compares] + · exact compare_lt_iff_lt + · exact compare_eq_iff_eq + · exact compare_gt_iff_gt + +theorem cmp_eq_compare (a b : α) : cmp a b = compare a b := by + refine ((compare_iff ..).2 ?_).symm + unfold cmp cmpUsing; split_ifs with h1 h2 + · exact h1 + · exact h2 + · exact le_antisymm (not_lt.1 h2) (not_lt.1 h1) + +theorem cmp_eq_compareOfLessAndEq (a b : α) : cmp a b = compareOfLessAndEq a b := + (cmp_eq_compare ..).trans (LinearOrder.compare_eq_compareOfLessAndEq ..) + +instance : Batteries.LawfulCmp (compare (α := α)) where + symm a b := by + cases h : compare a b <;> + simp only [Ordering.swap] <;> symm + · exact compare_gt_iff_gt.2 <| compare_lt_iff_lt.1 h + · exact compare_eq_iff_eq.2 <| compare_eq_iff_eq.1 h |>.symm + · exact compare_lt_iff_lt.2 <| compare_gt_iff_gt.1 h + le_trans := fun h₁ h₂ ↦ + compare_le_iff_le.2 <| le_trans (compare_le_iff_le.1 h₁) (compare_le_iff_le.1 h₂) + cmp_iff_beq := by simp [compare_eq_iff_eq] + cmp_iff_lt := by simp [compare_lt_iff_lt] + cmp_iff_le := by simp [compare_le_iff_le] + +end Ord + +end LinearOrder diff --git a/Mathlib/Order/Defs/PartialOrder.lean b/Mathlib/Order/Defs/PartialOrder.lean new file mode 100644 index 0000000000000..377677f6d9855 --- /dev/null +++ b/Mathlib/Order/Defs/PartialOrder.lean @@ -0,0 +1,158 @@ +/- +Copyright (c) 2016 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura +-/ +import Batteries.Tactic.Trans +import Mathlib.Tactic.ExtendDoc +import Mathlib.Tactic.Lemma +import Mathlib.Tactic.SplitIfs +import Mathlib.Tactic.TypeStar + +/-! +# Orders + +Defines classes for preorders and partial orders +and proves some basic lemmas about them. +-/ + +variable {α : Type*} + +section Preorder + +/-! +### Definition of `Preorder` and lemmas about types with a `Preorder` +-/ + +/-- A preorder is a reflexive, transitive relation `≤` with `a < b` defined in the obvious way. -/ +class Preorder (α : Type*) extends LE α, LT α where + le_refl : ∀ a : α, a ≤ a + le_trans : ∀ a b c : α, a ≤ b → b ≤ c → a ≤ c + lt := fun a b => a ≤ b ∧ ¬b ≤ a + lt_iff_le_not_le : ∀ a b : α, a < b ↔ a ≤ b ∧ ¬b ≤ a := by intros; rfl + +variable [Preorder α] {a b c : α} + +/-- The relation `≤` on a preorder is reflexive. -/ +@[refl, simp] lemma le_refl : ∀ a : α, a ≤ a := Preorder.le_refl + +/-- A version of `le_refl` where the argument is implicit -/ +lemma le_rfl : a ≤ a := le_refl a + +/-- The relation `≤` on a preorder is transitive. -/ +@[trans] lemma le_trans : a ≤ b → b ≤ c → a ≤ c := Preorder.le_trans _ _ _ + +lemma lt_iff_le_not_le : a < b ↔ a ≤ b ∧ ¬b ≤ a := Preorder.lt_iff_le_not_le _ _ + +lemma lt_of_le_not_le (hab : a ≤ b) (hba : ¬ b ≤ a) : a < b := lt_iff_le_not_le.2 ⟨hab, hba⟩ + +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] +theorem le_not_le_of_lt : ∀ {a b : α}, a < b → a ≤ b ∧ ¬b ≤ a + | _a, _b, hab => lt_iff_le_not_le.mp hab + +lemma le_of_eq (hab : a = b) : a ≤ b := by rw [hab] +lemma le_of_lt (hab : a < b) : a ≤ b := (lt_iff_le_not_le.1 hab).1 +lemma not_le_of_lt (hab : a < b) : ¬ b ≤ a := (lt_iff_le_not_le.1 hab).2 +lemma not_le_of_gt (hab : a > b) : ¬a ≤ b := not_le_of_lt hab +lemma not_lt_of_le (hab : a ≤ b) : ¬ b < a := imp_not_comm.1 not_le_of_lt hab +lemma not_lt_of_ge (hab : a ≥ b) : ¬a < b := not_lt_of_le hab + +alias LT.lt.not_le := not_le_of_lt +alias LE.le.not_lt := not_lt_of_le + +@[trans] lemma ge_trans : a ≥ b → b ≥ c → a ≥ c := fun h₁ h₂ => le_trans h₂ h₁ + +lemma lt_irrefl (a : α) : ¬a < a := fun h ↦ not_le_of_lt h le_rfl +lemma gt_irrefl (a : α) : ¬a > a := lt_irrefl _ + +@[trans] lemma lt_of_lt_of_le (hab : a < b) (hbc : b ≤ c) : a < c := + lt_of_le_not_le (le_trans (le_of_lt hab) hbc) fun hca ↦ not_le_of_lt hab (le_trans hbc hca) + +@[trans] lemma lt_of_le_of_lt (hab : a ≤ b) (hbc : b < c) : a < c := + lt_of_le_not_le (le_trans hab (le_of_lt hbc)) fun hca ↦ not_le_of_lt hbc (le_trans hca hab) + +@[trans] lemma gt_of_gt_of_ge (h₁ : a > b) (h₂ : b ≥ c) : a > c := lt_of_le_of_lt h₂ h₁ +@[trans] lemma gt_of_ge_of_gt (h₁ : a ≥ b) (h₂ : b > c) : a > c := lt_of_lt_of_le h₂ h₁ + +@[trans] lemma lt_trans (hab : a < b) (hbc : b < c) : a < c := lt_of_lt_of_le hab (le_of_lt hbc) +@[trans] lemma gt_trans : a > b → b > c → a > c := fun h₁ h₂ => lt_trans h₂ h₁ + +lemma ne_of_lt (h : a < b) : a ≠ b := fun he => absurd h (he ▸ lt_irrefl a) +lemma ne_of_gt (h : b < a) : a ≠ b := fun he => absurd h (he ▸ lt_irrefl a) +lemma lt_asymm (h : a < b) : ¬b < a := fun h1 : b < a => lt_irrefl a (lt_trans h h1) + +alias not_lt_of_gt := lt_asymm +alias not_lt_of_lt := lt_asymm + +lemma le_of_lt_or_eq (h : a < b ∨ a = b) : a ≤ b := h.elim le_of_lt le_of_eq +lemma le_of_eq_or_lt (h : a = b ∨ a < b) : a ≤ b := h.elim le_of_eq le_of_lt + +instance (priority := 900) : @Trans α α α LE.le LE.le LE.le := ⟨le_trans⟩ +instance (priority := 900) : @Trans α α α LT.lt LT.lt LT.lt := ⟨lt_trans⟩ +instance (priority := 900) : @Trans α α α LT.lt LE.le LT.lt := ⟨lt_of_lt_of_le⟩ +instance (priority := 900) : @Trans α α α LE.le LT.lt LT.lt := ⟨lt_of_le_of_lt⟩ +instance (priority := 900) : @Trans α α α GE.ge GE.ge GE.ge := ⟨ge_trans⟩ +instance (priority := 900) : @Trans α α α GT.gt GT.gt GT.gt := ⟨gt_trans⟩ +instance (priority := 900) : @Trans α α α GT.gt GE.ge GT.gt := ⟨gt_of_gt_of_ge⟩ +instance (priority := 900) : @Trans α α α GE.ge GT.gt GT.gt := ⟨gt_of_ge_of_gt⟩ + +/-- `<` is decidable if `≤` is. -/ +def decidableLTOfDecidableLE [@DecidableRel α (· ≤ ·)] : @DecidableRel α (· < ·) + | a, b => + if hab : a ≤ b then + if hba : b ≤ a then isFalse fun hab' => not_le_of_gt hab' hba + else isTrue <| lt_of_le_not_le hab hba + else isFalse fun hab' => hab (le_of_lt hab') + +end Preorder + +section PartialOrder + +/-! +### Definition of `PartialOrder` and lemmas about types with a partial order +-/ + +/-- A partial order is a reflexive, transitive, antisymmetric relation `≤`. -/ +class PartialOrder (α : Type*) extends Preorder α where + le_antisymm : ∀ a b : α, a ≤ b → b ≤ a → a = b + +variable [PartialOrder α] {a b : α} + +lemma le_antisymm : a ≤ b → b ≤ a → a = b := PartialOrder.le_antisymm _ _ + +alias eq_of_le_of_le := le_antisymm + +lemma le_antisymm_iff : a = b ↔ a ≤ b ∧ b ≤ a := + ⟨fun e => ⟨le_of_eq e, le_of_eq e.symm⟩, fun ⟨h1, h2⟩ => le_antisymm h1 h2⟩ + +lemma lt_of_le_of_ne : a ≤ b → a ≠ b → a < b := fun h₁ h₂ => + lt_of_le_not_le h₁ <| mt (le_antisymm h₁) h₂ + +/-- Equality is decidable if `≤` is. -/ +def decidableEqOfDecidableLE [@DecidableRel α (· ≤ ·)] : DecidableEq α + | a, b => + if hab : a ≤ b then + if hba : b ≤ a then isTrue (le_antisymm hab hba) else isFalse fun heq => hba (heq ▸ le_refl _) + else isFalse fun heq => hab (heq ▸ le_refl _) + +namespace Decidable + +variable [@DecidableRel α (· ≤ ·)] + +lemma lt_or_eq_of_le (hab : a ≤ b) : a < b ∨ a = b := + if hba : b ≤ a then Or.inr (le_antisymm hab hba) else Or.inl (lt_of_le_not_le hab hba) + +lemma eq_or_lt_of_le (hab : a ≤ b) : a = b ∨ a < b := + (lt_or_eq_of_le hab).symm + +lemma le_iff_lt_or_eq : a ≤ b ↔ a < b ∨ a = b := + ⟨lt_or_eq_of_le, le_of_lt_or_eq⟩ + +end Decidable + +attribute [local instance] Classical.propDecidable + +lemma lt_or_eq_of_le : a ≤ b → a < b ∨ a = b := Decidable.lt_or_eq_of_le +lemma le_iff_lt_or_eq : a ≤ b ↔ a < b ∨ a = b := Decidable.le_iff_lt_or_eq + +end PartialOrder diff --git a/Mathlib/Order/Defs/Unbundled.lean b/Mathlib/Order/Defs/Unbundled.lean new file mode 100644 index 0000000000000..c666d5f19cd3b --- /dev/null +++ b/Mathlib/Order/Defs/Unbundled.lean @@ -0,0 +1,278 @@ +/- +Copyright (c) 2016 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura +-/ +import Mathlib.Data.Set.Defs +import Mathlib.Tactic.ExtendDoc +import Mathlib.Tactic.Lemma +import Mathlib.Tactic.SplitIfs +import Mathlib.Tactic.TypeStar + +/-! +# Orders + +Defines classes for preorders, partial orders, and linear orders +and proves some basic lemmas about them. +-/ + +/-! ### Unbundled classes -/ + +/-- An empty relation does not relate any elements. -/ +@[nolint unusedArguments] def EmptyRelation {α : Sort*} := fun _ _ : α ↦ False + +/-- `IsIrrefl X r` means the binary relation `r` on `X` is irreflexive (that is, `r x x` never +holds). -/ +class IsIrrefl (α : Sort*) (r : α → α → Prop) : Prop where + irrefl : ∀ a, ¬r a a + +/-- `IsRefl X r` means the binary relation `r` on `X` is reflexive. -/ +class IsRefl (α : Sort*) (r : α → α → Prop) : Prop where + refl : ∀ a, r a a + +/-- `IsSymm X r` means the binary relation `r` on `X` is symmetric. -/ +class IsSymm (α : Sort*) (r : α → α → Prop) : Prop where + symm : ∀ a b, r a b → r b a + +/-- `IsAsymm X r` means that the binary relation `r` on `X` is asymmetric, that is, +`r a b → ¬ r b a`. -/ +class IsAsymm (α : Sort*) (r : α → α → Prop) : Prop where + asymm : ∀ a b, r a b → ¬r b a + +/-- `IsAntisymm X r` means the binary relation `r` on `X` is antisymmetric. -/ +class IsAntisymm (α : Sort*) (r : α → α → Prop) : Prop where + antisymm : ∀ a b, r a b → r b a → a = b + +instance (priority := 100) IsAsymm.toIsAntisymm {α : Sort*} (r : α → α → Prop) [IsAsymm α r] : + IsAntisymm α r where + antisymm _ _ hx hy := (IsAsymm.asymm _ _ hx hy).elim + +/-- `IsTrans X r` means the binary relation `r` on `X` is transitive. -/ +class IsTrans (α : Sort*) (r : α → α → Prop) : Prop where + trans : ∀ a b c, r a b → r b c → r a c + +instance {α : Sort*} {r : α → α → Prop} [IsTrans α r] : Trans r r r := + ⟨IsTrans.trans _ _ _⟩ + +instance (priority := 100) {α : Sort*} {r : α → α → Prop} [Trans r r r] : IsTrans α r := + ⟨fun _ _ _ => Trans.trans⟩ + +/-- `IsTotal X r` means that the binary relation `r` on `X` is total, that is, that for any +`x y : X` we have `r x y` or `r y x`. -/ +class IsTotal (α : Sort*) (r : α → α → Prop) : Prop where + total : ∀ a b, r a b ∨ r b a + +/-- `IsPreorder X r` means that the binary relation `r` on `X` is a pre-order, that is, reflexive +and transitive. -/ +class IsPreorder (α : Sort*) (r : α → α → Prop) extends IsRefl α r, IsTrans α r : Prop + +/-- `IsPartialOrder X r` means that the binary relation `r` on `X` is a partial order, that is, +`IsPreorder X r` and `IsAntisymm X r`. -/ +class IsPartialOrder (α : Sort*) (r : α → α → Prop) extends IsPreorder α r, IsAntisymm α r : Prop + +/-- `IsLinearOrder X r` means that the binary relation `r` on `X` is a linear order, that is, +`IsPartialOrder X r` and `IsTotal X r`. -/ +class IsLinearOrder (α : Sort*) (r : α → α → Prop) extends IsPartialOrder α r, IsTotal α r : Prop + +/-- `IsEquiv X r` means that the binary relation `r` on `X` is an equivalence relation, that +is, `IsPreorder X r` and `IsSymm X r`. -/ +class IsEquiv (α : Sort*) (r : α → α → Prop) extends IsPreorder α r, IsSymm α r : Prop + +/-- `IsStrictOrder X r` means that the binary relation `r` on `X` is a strict order, that is, +`IsIrrefl X r` and `IsTrans X r`. -/ +class IsStrictOrder (α : Sort*) (r : α → α → Prop) extends IsIrrefl α r, IsTrans α r : Prop + +/-- `IsStrictWeakOrder X lt` means that the binary relation `lt` on `X` is a strict weak order, +that is, `IsStrictOrder X lt` and `¬lt a b ∧ ¬lt b a → ¬lt b c ∧ ¬lt c b → ¬lt a c ∧ ¬lt c a`. -/ +class IsStrictWeakOrder (α : Sort*) (lt : α → α → Prop) extends IsStrictOrder α lt : Prop where + incomp_trans : ∀ a b c, ¬lt a b ∧ ¬lt b a → ¬lt b c ∧ ¬lt c b → ¬lt a c ∧ ¬lt c a + +/-- `IsTrichotomous X lt` means that the binary relation `lt` on `X` is trichotomous, that is, +either `lt a b` or `a = b` or `lt b a` for any `a` and `b`. -/ +class IsTrichotomous (α : Sort*) (lt : α → α → Prop) : Prop where + trichotomous : ∀ a b, lt a b ∨ a = b ∨ lt b a + +/-- `IsStrictTotalOrder X lt` means that the binary relation `lt` on `X` is a strict total order, +that is, `IsTrichotomous X lt` and `IsStrictOrder X lt`. -/ +class IsStrictTotalOrder (α : Sort*) (lt : α → α → Prop) extends IsTrichotomous α lt, + IsStrictOrder α lt : Prop + +/-- Equality is an equivalence relation. -/ +instance eq_isEquiv (α : Sort*) : IsEquiv α (· = ·) where + symm := @Eq.symm _ + trans := @Eq.trans _ + refl := Eq.refl + +/-- `Iff` is an equivalence relation. -/ +instance iff_isEquiv : IsEquiv Prop Iff where + symm := @Iff.symm + trans := @Iff.trans + refl := @Iff.refl + +section + +variable {α : Sort*} {r : α → α → Prop} {a b c : α} + +/-- Local notation for an arbitrary binary relation `r`. -/ +local infixl:50 " ≺ " => r + +lemma irrefl [IsIrrefl α r] (a : α) : ¬a ≺ a := IsIrrefl.irrefl a +lemma refl [IsRefl α r] (a : α) : a ≺ a := IsRefl.refl a +lemma trans [IsTrans α r] : a ≺ b → b ≺ c → a ≺ c := IsTrans.trans _ _ _ +lemma symm [IsSymm α r] : a ≺ b → b ≺ a := IsSymm.symm _ _ +lemma antisymm [IsAntisymm α r] : a ≺ b → b ≺ a → a = b := IsAntisymm.antisymm _ _ +lemma asymm [IsAsymm α r] : a ≺ b → ¬b ≺ a := IsAsymm.asymm _ _ + +lemma trichotomous [IsTrichotomous α r] : ∀ a b : α, a ≺ b ∨ a = b ∨ b ≺ a := + IsTrichotomous.trichotomous + +instance (priority := 90) isAsymm_of_isTrans_of_isIrrefl [IsTrans α r] [IsIrrefl α r] : + IsAsymm α r := + ⟨fun a _b h₁ h₂ => absurd (_root_.trans h₁ h₂) (irrefl a)⟩ + +instance IsIrrefl.decide [DecidableRel r] [IsIrrefl α r] : + IsIrrefl α (fun a b => decide (r a b) = true) where + irrefl := fun a => by simpa using irrefl a + +instance IsRefl.decide [DecidableRel r] [IsRefl α r] : + IsRefl α (fun a b => decide (r a b) = true) where + refl := fun a => by simpa using refl a + +instance IsTrans.decide [DecidableRel r] [IsTrans α r] : + IsTrans α (fun a b => decide (r a b) = true) where + trans := fun a b c => by simpa using trans a b c + +instance IsSymm.decide [DecidableRel r] [IsSymm α r] : + IsSymm α (fun a b => decide (r a b) = true) where + symm := fun a b => by simpa using symm a b + +instance IsAntisymm.decide [DecidableRel r] [IsAntisymm α r] : + IsAntisymm α (fun a b => decide (r a b) = true) where + antisymm := fun a b h₁ h₂ => antisymm _ _ (by simpa using h₁) (by simpa using h₂) + +instance IsAsymm.decide [DecidableRel r] [IsAsymm α r] : + IsAsymm α (fun a b => decide (r a b) = true) where + asymm := fun a b => by simpa using asymm a b + +instance IsTotal.decide [DecidableRel r] [IsTotal α r] : + IsTotal α (fun a b => decide (r a b) = true) where + total := fun a b => by simpa using total a b + +instance IsTrichotomous.decide [DecidableRel r] [IsTrichotomous α r] : + IsTrichotomous α (fun a b => decide (r a b) = true) where + trichotomous := fun a b => by simpa using trichotomous a b + +variable (r) + +@[elab_without_expected_type] lemma irrefl_of [IsIrrefl α r] (a : α) : ¬a ≺ a := irrefl a +@[elab_without_expected_type] lemma refl_of [IsRefl α r] (a : α) : a ≺ a := refl a +@[elab_without_expected_type] lemma trans_of [IsTrans α r] : a ≺ b → b ≺ c → a ≺ c := _root_.trans +@[elab_without_expected_type] lemma symm_of [IsSymm α r] : a ≺ b → b ≺ a := symm +@[elab_without_expected_type] lemma asymm_of [IsAsymm α r] : a ≺ b → ¬b ≺ a := asymm + +@[elab_without_expected_type] +lemma total_of [IsTotal α r] (a b : α) : a ≺ b ∨ b ≺ a := IsTotal.total _ _ + +@[elab_without_expected_type] +lemma trichotomous_of [IsTrichotomous α r] : ∀ a b : α, a ≺ b ∨ a = b ∨ b ≺ a := trichotomous + +section + +/-- `IsRefl` as a definition, suitable for use in proofs. -/ +def Reflexive := ∀ x, x ≺ x + +/-- `IsSymm` as a definition, suitable for use in proofs. -/ +def Symmetric := ∀ ⦃x y⦄, x ≺ y → y ≺ x + +/-- `IsTrans` as a definition, suitable for use in proofs. -/ +def Transitive := ∀ ⦃x y z⦄, x ≺ y → y ≺ z → x ≺ z + +/-- `IsIrrefl` as a definition, suitable for use in proofs. -/ +def Irreflexive := ∀ x, ¬x ≺ x + +/-- `IsAntisymm` as a definition, suitable for use in proofs. -/ +def AntiSymmetric := ∀ ⦃x y⦄, x ≺ y → y ≺ x → x = y + +/-- `IsTotal` as a definition, suitable for use in proofs. -/ +def Total := ∀ x y, x ≺ y ∨ y ≺ x + +@[deprecated Equivalence.refl (since := "2024-09-13")] +theorem Equivalence.reflexive (h : Equivalence r) : Reflexive r := h.refl + +@[deprecated Equivalence.symm (since := "2024-09-13")] +theorem Equivalence.symmetric (h : Equivalence r) : Symmetric r := + fun _ _ ↦ h.symm + +@[deprecated Equivalence.trans (since := "2024-09-13")] +theorem Equivalence.transitive (h : Equivalence r) : Transitive r := + fun _ _ _ ↦ h.trans + +variable {β : Sort*} (r : β → β → Prop) (f : α → β) + +@[deprecated "No deprecation message was provided." (since := "2024-09-13")] +theorem InvImage.trans (h : Transitive r) : Transitive (InvImage r f) := + fun (a₁ a₂ a₃ : α) (h₁ : InvImage r f a₁ a₂) (h₂ : InvImage r f a₂ a₃) ↦ h h₁ h₂ + +@[deprecated "No deprecation message was provided." (since := "2024-09-13")] +theorem InvImage.irreflexive (h : Irreflexive r) : Irreflexive (InvImage r f) := + fun (a : α) (h₁ : InvImage r f a a) ↦ h (f a) h₁ + +end + +end + +/-! ### Minimal and maximal -/ + +section LE + +variable {α : Type*} [LE α] {P : α → Prop} {x y : α} + +/-- `Minimal P x` means that `x` is a minimal element satisfying `P`. -/ +def Minimal (P : α → Prop) (x : α) : Prop := P x ∧ ∀ ⦃y⦄, P y → y ≤ x → x ≤ y + +/-- `Maximal P x` means that `x` is a maximal element satisfying `P`. -/ +def Maximal (P : α → Prop) (x : α) : Prop := P x ∧ ∀ ⦃y⦄, P y → x ≤ y → y ≤ x + +lemma Minimal.prop (h : Minimal P x) : P x := + h.1 + +lemma Maximal.prop (h : Maximal P x) : P x := + h.1 + +lemma Minimal.le_of_le (h : Minimal P x) (hy : P y) (hle : y ≤ x) : x ≤ y := + h.2 hy hle + +lemma Maximal.le_of_ge (h : Maximal P x) (hy : P y) (hge : x ≤ y) : y ≤ x := + h.2 hy hge + +end LE + +/-! ### Upper and lower sets -/ + +/-- An upper set in an order `α` is a set such that any element greater than one of its members is +also a member. Also called up-set, upward-closed set. -/ +def IsUpperSet {α : Type*} [LE α] (s : Set α) : Prop := + ∀ ⦃a b : α⦄, a ≤ b → a ∈ s → b ∈ s + +/-- A lower set in an order `α` is a set such that any element less than one of its members is also +a member. Also called down-set, downward-closed set. -/ +def IsLowerSet {α : Type*} [LE α] (s : Set α) : Prop := + ∀ ⦃a b : α⦄, b ≤ a → a ∈ s → b ∈ s + +@[inherit_doc IsUpperSet] +structure UpperSet (α : Type*) [LE α] where + /-- The carrier of an `UpperSet`. -/ + carrier : Set α + /-- The carrier of an `UpperSet` is an upper set. -/ + upper' : IsUpperSet carrier + +extend_docs UpperSet before "The type of upper sets of an order." + +@[inherit_doc IsLowerSet] +structure LowerSet (α : Type*) [LE α] where + /-- The carrier of a `LowerSet`. -/ + carrier : Set α + /-- The carrier of a `LowerSet` is a lower set. -/ + lower' : IsLowerSet carrier + +extend_docs LowerSet before "The type of lower sets of an order." diff --git a/Mathlib/Order/Interval/Set/Defs.lean b/Mathlib/Order/Interval/Set/Defs.lean index a233ef08b886c..1ca919215430a 100644 --- a/Mathlib/Order/Interval/Set/Defs.lean +++ b/Mathlib/Order/Interval/Set/Defs.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ import Mathlib.Data.Set.Defs -import Mathlib.Order.Defs +import Mathlib.Order.Defs.PartialOrder /-! # Intervals diff --git a/Mathlib/Tactic/GCongr/Core.lean b/Mathlib/Tactic/GCongr/Core.lean index 90aeed7ba3815..f225f2c67bf80 100644 --- a/Mathlib/Tactic/GCongr/Core.lean +++ b/Mathlib/Tactic/GCongr/Core.lean @@ -4,11 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Heather Macbeth -/ import Lean -import Mathlib.Order.Defs -import Mathlib.Tactic.Core -import Mathlib.Tactic.GCongr.ForwardAttr import Batteries.Lean.Except import Batteries.Tactic.Exact +import Mathlib.Tactic.Core +import Mathlib.Tactic.GCongr.ForwardAttr +import Mathlib.Order.Defs.PartialOrder /-! # The `gcongr` ("generalized congruence") tactic diff --git a/Mathlib/Tactic/PushNeg.lean b/Mathlib/Tactic/PushNeg.lean index 07699624b4fc5..39b7ff1d2cd6b 100644 --- a/Mathlib/Tactic/PushNeg.lean +++ b/Mathlib/Tactic/PushNeg.lean @@ -7,7 +7,7 @@ Authors: Patrick Massot, Simon Hudon, Alice Laroche, Frédéric Dupuis, Jireh Lo import Lean.Elab.Tactic.Location import Mathlib.Data.Set.Defs import Mathlib.Logic.Basic -import Mathlib.Order.Defs +import Mathlib.Order.Defs.LinearOrder import Mathlib.Tactic.Conv /-! diff --git a/MathlibTest/push_neg.lean b/MathlibTest/push_neg.lean index 571b3a5f85cab..ccfb7e34cc4a6 100644 --- a/MathlibTest/push_neg.lean +++ b/MathlibTest/push_neg.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Alice Laroche, Frédéric Dupuis, Jireh Loreaux -/ -import Mathlib.Order.Defs +import Mathlib.Order.Defs.LinearOrder import Mathlib.Tactic.PushNeg private axiom test_sorry : ∀ {α}, α diff --git a/scripts/noshake.json b/scripts/noshake.json index b3c884f2347c8..2ca0558661bde 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -250,7 +250,8 @@ "Mathlib.Tactic.IrreducibleDef": ["Mathlib.Data.Subtype"], "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.Core": + ["Mathlib.Order.Defs", "Mathlib.Order.Defs.PartialOrder"], "Mathlib.Tactic.GCongr": ["Mathlib.Tactic.Positivity.Core"], "Mathlib.Tactic.FunProp.RefinedDiscrTree": ["Mathlib.Algebra.Group.Pi.Basic"], "Mathlib.Tactic.FunProp.Mor": ["Mathlib.Data.FunLike.Basic"], @@ -336,6 +337,7 @@ "Mathlib.Order.RelClasses": ["Batteries.WF"], "Mathlib.Order.Filter.ListTraverse": ["Mathlib.Control.Traversable.Instances"], + "Mathlib.Order.Defs.PartialOrder": ["Batteries.Tactic.Trans"], "Mathlib.Order.Defs": ["Batteries.Tactic.Trans"], "Mathlib.Order.BoundedOrder.Basic": ["Mathlib.Tactic.Finiteness.Attr"], "Mathlib.Order.BoundedOrder": ["Mathlib.Tactic.Finiteness.Attr"], @@ -349,6 +351,7 @@ "Mathlib.Logic.Function.Defs": ["Mathlib.Tactic.AdaptationNote"], "Mathlib.Logic.Function.Basic": ["Batteries.Tactic.Init"], "Mathlib.Logic.Equiv.Defs": ["Mathlib.Data.Bool.Basic"], + "Mathlib.Logic.Basic": ["Batteries.Tactic.Trans"], "Mathlib.LinearAlgebra.Matrix.Transvection": ["Mathlib.Data.Matrix.DMatrix"], "Mathlib.LinearAlgebra.DFinsupp": ["Mathlib.LinearAlgebra.Finsupp.SumProd"], "Mathlib.LinearAlgebra.Basic": ["Mathlib.Algebra.BigOperators.Group.Finset"], From 841cc90f3826f488626a68e34117d45f83cff8e5 Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Tue, 26 Nov 2024 11:45:38 +0000 Subject: [PATCH 315/829] feat(Algebra/Category): API for `Under R` (#19058) We provide API for `Under R` where `R : CommRingCat`. We also add some `simp` lemmas for the various ring categories. --- Mathlib.lean | 1 + .../ModuleCat/Differentials/Presheaf.lean | 2 +- Mathlib/Algebra/Category/Ring/Basic.lean | 39 +++- .../Algebra/Category/Ring/Constructions.lean | 11 ++ .../Algebra/Category/Ring/Under/Basic.lean | 178 ++++++++++++++++++ Mathlib/CategoryTheory/Comma/Over.lean | 20 ++ Mathlib/RingTheory/TensorProduct/Basic.lean | 9 + 7 files changed, 254 insertions(+), 6 deletions(-) create mode 100644 Mathlib/Algebra/Category/Ring/Under/Basic.lean diff --git a/Mathlib.lean b/Mathlib.lean index c81570dc36eb0..fb9f9c436a030 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -152,6 +152,7 @@ import Mathlib.Algebra.Category.Ring.FilteredColimits import Mathlib.Algebra.Category.Ring.Instances import Mathlib.Algebra.Category.Ring.Limits import Mathlib.Algebra.Category.Ring.LinearAlgebra +import Mathlib.Algebra.Category.Ring.Under.Basic import Mathlib.Algebra.Category.Semigrp.Basic import Mathlib.Algebra.Central.Basic import Mathlib.Algebra.Central.Defs diff --git a/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean b/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean index eef6ac3891d20..5ea1401c768c2 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean @@ -183,7 +183,7 @@ noncomputable def relativeDifferentials' : obj X := CommRingCat.KaehlerDifferential (φ'.app X) map f := CommRingCat.KaehlerDifferential.map (φ'.naturality f) map_id _ := by ext; simp; rfl - map_comp _ _ := by ext; simp; rfl + map_comp _ _ := by ext; simp attribute [simp] relativeDifferentials'_obj diff --git a/Mathlib/Algebra/Category/Ring/Basic.lean b/Mathlib/Algebra/Category/Ring/Basic.lean index fe80c9b8c463c..61b53a227b0c6 100644 --- a/Mathlib/Algebra/Category/Ring/Basic.lean +++ b/Mathlib/Algebra/Category/Ring/Basic.lean @@ -98,6 +98,14 @@ theorem coe_of (R : Type u) [Semiring R] : (SemiRingCat.of R : Type u) = R := @DFunLike.coe (X →+* Z) X (fun _ ↦ Z) _ (RingHom.comp g f) := rfl +/-- Variant of `SemiRingCat.coe_comp_of` for morphisms. -/ +@[simp] theorem coe_comp_of' {X Z : Type u} {Y : SemiRingCat.{u}} + [Semiring X] [Semiring Z] + (f : of X ⟶ Y) (g : Y ⟶ of Z) : + @DFunLike.coe no_index (of X ⟶ of Z) X (fun _ ↦ Z) _ (f ≫ g) = + @DFunLike.coe (X →+* Z) X (fun _ ↦ Z) _ (g.comp f) := by + rfl + -- Sometimes neither the `ext` lemma for `SemiRingCat` nor for `RingHom` is applicable, -- because of incomplete unfolding of `SemiRingCat.of X ⟶ SemiRingCat.of Y := X →+* Y`, -- but this one will fire. @@ -142,7 +150,7 @@ theorem ofHom_apply {R S : Type u} [Semiring R] [Semiring S] (f : R →+* S) (x /-- A variant of `ofHom_apply` that makes `simpNF` happy -/ @[simp] theorem ofHom_apply' {R S : Type u} [Semiring R] [Semiring S] (f : R →+* S) (x : R) : - DFunLike.coe (α := R) (β := fun _ ↦ S) (ofHom f) x = f x := rfl + @DFunLike.coe no_index _ R (fun _ ↦ S) _ (ofHom f) x = f x := rfl /-- Ring equivalence are isomorphisms in category of semirings @@ -221,7 +229,7 @@ theorem coe_of (R : Type u) [Ring R] : (RingCat.of R : Type u) = R := /-- A variant of `ofHom_apply` that makes `simpNF` happy -/ @[simp] theorem ofHom_apply' {R S : Type u} [Ring R] [Ring S] (f : R →+* S) (x : R) : - DFunLike.coe (α := R) (β := fun _ ↦ S) (ofHom f) x = f x := rfl + @DFunLike.coe no_index _ R (fun _ ↦ S) _ (ofHom f) x = f x := rfl -- Coercing the identity morphism, as a ring homomorphism, gives the identity function. @[simp] theorem coe_ringHom_id {X : RingCat} : @@ -247,6 +255,13 @@ theorem ofHom_apply' {R S : Type u} [Ring R] [Ring S] (f : R →+* S) (x : R) : @DFunLike.coe (X →+* Z) X (fun _ ↦ Z) _ (RingHom.comp g f) := rfl +/-- Variant of `RingCat.coe_comp_of` for morphisms. -/ +@[simp] theorem coe_comp_of' {X Z : Type u} {Y : RingCat.{u}} [Ring X] [Ring Z] + (f : of X ⟶ Y) (g : Y ⟶ of Z) : + @DFunLike.coe no_index (of X ⟶ of Z) X (fun _ ↦ Z) _ (f ≫ g) = + @DFunLike.coe (X →+* Z) X (fun _ ↦ Z) _ (g.comp f) := by + rfl + -- Sometimes neither the `ext` lemma for `RingCat` nor for `RingHom` is applicable, -- because of incomplete unfolding of `RingCat.of X ⟶ RingCat.of Y := X →+* Y`, -- but this one will fire. @@ -340,7 +355,7 @@ theorem ofHom_apply {R S : Type u} [CommSemiring R] [CommSemiring S] (f : R →+ /-- A variant of `ofHom_apply` that makes `simpNF` happy -/ @[simp] theorem ofHom_apply' {R S : Type u} [CommSemiring R] [CommSemiring S] (f : R →+* S) (x : R) : - DFunLike.coe (α := R) (β := fun _ ↦ S) (ofHom f) x = f x := rfl + @DFunLike.coe no_index _ R (fun _ ↦ S) _ (ofHom f) x = f x := rfl instance : Inhabited CommSemiRingCat := ⟨of PUnit⟩ @@ -376,6 +391,14 @@ theorem coe_of (R : Type u) [CommSemiring R] : (CommSemiRingCat.of R : Type u) = @DFunLike.coe (X →+* Z) X (fun _ ↦ Z) _ (RingHom.comp g f) := rfl +/-- Variant of `CommSemiRingCat.coe_comp_of` for morphisms. -/ +@[simp] theorem coe_comp_of' {X Z : Type u} {Y : CommSemiRingCat.{u}} + [CommSemiring X] [CommSemiring Z] + (f : of X ⟶ Y) (g : Y ⟶ of Z) : + @DFunLike.coe no_index (of X ⟶ of Z) X (fun _ ↦ Z) _ (f ≫ g) = + @DFunLike.coe (X →+* Z) X (fun _ ↦ Z) _ (g.comp f) := by + rfl + -- Sometimes neither the `ext` lemma for `CommSemiRingCat` nor for `RingHom` is applicable, -- because of incomplete unfolding of `CommSemiRingCat.of X ⟶ CommSemiRingCat.of Y := X →+* Y`, -- but this one will fire. @@ -491,8 +514,7 @@ theorem ofHom_apply {R S : Type u} [CommRing R] [CommRing S] (f : R →+* S) (x /-- A variant of `ofHom_apply` that makes `simpNF` happy -/ @[simp] theorem ofHom_apply' {R S : Type u} [CommRing R] [CommRing S] (f : R →+* S) (x : R) : - DFunLike.coe (α := R) (β := fun _ ↦ S) (ofHom f) x = f x := rfl - + @DFunLike.coe no_index _ R (fun _ ↦ S) _ (ofHom f) x = f x := rfl instance : Inhabited CommRingCat := ⟨of PUnit⟩ @@ -528,6 +550,13 @@ theorem coe_of (R : Type u) [CommRing R] : (CommRingCat.of R : Type u) = R := @DFunLike.coe (X →+* Z) X (fun _ ↦ Z) _ (RingHom.comp g f) := rfl +/-- Variant of `CommRingCat.coe_comp_of` for morphisms. -/ +@[simp] theorem coe_comp_of' {X Z : Type u} {Y : CommRingCat.{u}} [CommRing X] [CommRing Z] + (f : of X ⟶ Y) (g : Y ⟶ of Z) : + @DFunLike.coe no_index (of X ⟶ of Z) X (fun _ ↦ Z) _ (f ≫ g) = + @DFunLike.coe (X →+* Z) X (fun _ ↦ Z) _ (g.comp f) := by + rfl + -- Sometimes neither the `ext` lemma for `CommRingCat` nor for `RingHom` is applicable, -- because of incomplete unfolding of `CommRingCat.of X ⟶ CommRingCat.of Y := X →+* Y`, -- but this one will fire. diff --git a/Mathlib/Algebra/Category/Ring/Constructions.lean b/Mathlib/Algebra/Category/Ring/Constructions.lean index fea4b951aa7ee..ca21d51022790 100644 --- a/Mathlib/Algebra/Category/Ring/Constructions.lean +++ b/Mathlib/Algebra/Category/Ring/Constructions.lean @@ -5,6 +5,7 @@ Authors: Andrew Yang -/ import Mathlib.Algebra.Category.Ring.Instances import Mathlib.Algebra.Category.Ring.Limits +import Mathlib.CategoryTheory.Limits.Shapes.Pullback.CommSq import Mathlib.CategoryTheory.Limits.Shapes.StrictInitial import Mathlib.RingTheory.TensorProduct.Basic @@ -102,6 +103,16 @@ def pushoutCoconeIsColimit : Limits.IsColimit (pushoutCocone R A B) := rw [← h.map_mul, Algebra.TensorProduct.tmul_mul_tmul, mul_one, one_mul] rfl +lemma isPushout_tensorProduct (R A B : Type u) [CommRing R] [CommRing A] [CommRing B] + [Algebra R A] [Algebra R B] : + IsPushout (ofHom <| algebraMap R A) (ofHom <| algebraMap R B) + (ofHom <| Algebra.TensorProduct.includeLeftRingHom (R := R)) + (ofHom <| Algebra.TensorProduct.includeRight.toRingHom (R := R)) where + w := by + ext + simp + isColimit' := ⟨pushoutCoconeIsColimit R A B⟩ + end Pushout section Terminal diff --git a/Mathlib/Algebra/Category/Ring/Under/Basic.lean b/Mathlib/Algebra/Category/Ring/Under/Basic.lean new file mode 100644 index 0000000000000..a5e50b4d3c16b --- /dev/null +++ b/Mathlib/Algebra/Category/Ring/Under/Basic.lean @@ -0,0 +1,178 @@ +/- +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.Algebra.Category.Ring.Colimits +import Mathlib.Algebra.Category.Ring.Constructions +import Mathlib.CategoryTheory.Adjunction.Over + +/-! +# Under `CommRingCat` + +In this file we provide basic API for `Under R` when `R : CommRingCat`. `Under R` is +(equivalent to) the category of commutative `R`-algebras. For not necessarily commutative +algebras, use `AlgebraCat R` instead. +-/ + +noncomputable section + +universe u + +open TensorProduct CategoryTheory Limits + +variable {R S : CommRingCat.{u}} + +namespace CommRingCat + +instance : CoeSort (Under R) (Type u) where + coe A := A.right + +instance (A : Under R) : Algebra R A := RingHom.toAlgebra A.hom + +/-- Turn a morphism in `Under R` into an algebra homomorphism. -/ +def toAlgHom {A B : Under R} (f : A ⟶ B) : A →ₐ[R] B where + __ := f.right + commutes' a := by + have : (A.hom ≫ f.right) a = B.hom a := by simp + simpa only [Functor.const_obj_obj, Functor.id_obj, CommRingCat.comp_apply] using this + +@[simp] +lemma toAlgHom_id (A : Under R) : toAlgHom (𝟙 A) = AlgHom.id R A := rfl + +@[simp] +lemma toAlgHom_comp {A B C : Under R} (f : A ⟶ B) (g : B ⟶ C) : + toAlgHom (f ≫ g) = (toAlgHom g).comp (toAlgHom f) := rfl + +@[simp] +lemma toAlgHom_apply {A B : Under R} (f : A ⟶ B) (a : A) : + toAlgHom f a = f.right a := + rfl + +variable (R) in +/-- Make an object of `Under R` from an `R`-algebra. -/ +@[simps! hom, simps! (config := .lemmasOnly) right] +def mkUnder (A : Type u) [CommRing A] [Algebra R A] : Under R := + Under.mk (CommRingCat.ofHom <| algebraMap R A) + +@[ext] +lemma mkUnder_ext {A : Type u} [CommRing A] [Algebra R A] {B : Under R} + {f g : mkUnder R A ⟶ B} (h : ∀ a : A, f.right a = g.right a) : + f = g := by + ext x + exact h x + +end CommRingCat + +namespace AlgHom + +/-- Make a morphism in `Under R` from an algebra map. -/ +def toUnder {A B : Type u} [CommRing A] [CommRing B] [Algebra R A] [Algebra R B] + (f : A →ₐ[R] B) : CommRingCat.mkUnder R A ⟶ CommRingCat.mkUnder R B := + Under.homMk f.toRingHom <| by + ext a + exact f.commutes' a + +@[simp] +lemma toUnder_right {A B : Type u} [CommRing A] [CommRing B] [Algebra R A] + [Algebra R B] (f : A →ₐ[R] B) (a : A) : + f.toUnder.right a = f a := + rfl + +@[simp] +lemma toUnder_comp {A B C : Type u} [CommRing A] [CommRing B] [CommRing C] + [Algebra R A] [Algebra R B] [Algebra R C] (f : A →ₐ[R] B) (g : B →ₐ[R] C) : + (g.comp f).toUnder = f.toUnder ≫ g.toUnder := + rfl + +end AlgHom + +namespace AlgEquiv + +/-- Make an isomorphism in `Under R` from an algebra isomorphism. -/ +def toUnder {A B : Type u} [CommRing A] [CommRing B] [Algebra R A] [Algebra R B] + (f : A ≃ₐ[R] B) : + CommRingCat.mkUnder R A ≅ CommRingCat.mkUnder R B where + hom := f.toAlgHom.toUnder + inv := f.symm.toAlgHom.toUnder + hom_inv_id := by + ext (a : (CommRingCat.mkUnder R A).right) + simp + inv_hom_id := by + ext a + simp + +@[simp] +lemma toUnder_hom_right_apply {A B : Type u} [CommRing A] [CommRing B] [Algebra R A] + [Algebra R B] (f : A ≃ₐ[R] B) (a : A) : + f.toUnder.hom.right a = f a := rfl + +@[simp] +lemma toUnder_inv_right_apply {A B : Type u} [CommRing A] [CommRing B] [Algebra R A] + [Algebra R B] (f : A ≃ₐ[R] B) (b : B) : + f.toUnder.inv.right b = f.symm b := rfl + +@[simp] +lemma toUnder_trans {A B C : Type u} [CommRing A] [CommRing B] [CommRing C] + [Algebra R A] [Algebra R B] [Algebra R C] (f : A ≃ₐ[R] B) (g : B ≃ₐ[R] C) : + (f.trans g).toUnder = f.toUnder ≪≫ g.toUnder := + rfl + +end AlgEquiv + +namespace CommRingCat + +variable [Algebra R S] + +variable (R S) in +/-- The base change functor `A ↦ S ⊗[R] A`. -/ +@[simps! map_right] +def tensorProd : Under R ⥤ Under S where + obj A := mkUnder S (S ⊗[R] A) + map f := Algebra.TensorProduct.map (AlgHom.id S S) (toAlgHom f) |>.toUnder + map_comp {X Y Z} f g := by simp [Algebra.TensorProduct.map_id_comp] + +variable (S) in +/-- The natural isomorphism `S ⊗[R] A ≅ pushout A.hom (algebraMap R S)` in `Under S`. -/ +def tensorProdObjIsoPushoutObj (A : Under R) : + mkUnder S (S ⊗[R] A) ≅ (Under.pushout (algebraMap R S)).obj A := + Under.isoMk (CommRingCat.isPushout_tensorProduct R S A).flip.isoPushout <| by + simp only [Functor.const_obj_obj, Under.pushout_obj, Functor.id_obj, Under.mk_right, + mkUnder_hom, AlgHom.toRingHom_eq_coe, IsPushout.inr_isoPushout_hom, Under.mk_hom] + rfl + +@[reassoc (attr := simp)] +lemma pushout_inl_tensorProdObjIsoPushoutObj_inv_right (A : Under R) : + pushout.inl A.hom (algebraMap R S) ≫ (tensorProdObjIsoPushoutObj S A).inv.right = + (CommRingCat.ofHom <| Algebra.TensorProduct.includeRight.toRingHom) := by + simp [tensorProdObjIsoPushoutObj] + +@[reassoc (attr := simp)] +lemma pushout_inr_tensorProdObjIsoPushoutObj_inv_right (A : Under R) : + pushout.inr A.hom (algebraMap R S) ≫ + (tensorProdObjIsoPushoutObj S A).inv.right = + (CommRingCat.ofHom <| Algebra.TensorProduct.includeLeftRingHom) := by + simp [tensorProdObjIsoPushoutObj] + +variable (R S) in +/-- `A ↦ S ⊗[R] A` is naturally isomorphic to `A ↦ pushout A.hom (algebraMap R S)`. -/ +def tensorProdIsoPushout : tensorProd R S ≅ Under.pushout (algebraMap R S) := + NatIso.ofComponents (fun A ↦ tensorProdObjIsoPushoutObj S A) <| by + intro A B f + dsimp + rw [← cancel_epi (tensorProdObjIsoPushoutObj S A).inv] + ext : 1 + apply pushout.hom_ext + · rw [← cancel_mono (tensorProdObjIsoPushoutObj S B).inv.right] + ext x + simp [mkUnder_right] + · rw [← cancel_mono (tensorProdObjIsoPushoutObj S B).inv.right] + ext (x : S) + simp [mkUnder_right] + +@[simp] +lemma tensorProdIsoPushout_app (A : Under R) : + (tensorProdIsoPushout R S).app A = tensorProdObjIsoPushoutObj S A := + rfl + +end CommRingCat diff --git a/Mathlib/CategoryTheory/Comma/Over.lean b/Mathlib/CategoryTheory/Comma/Over.lean index 305664eb69cad..9f9609c579450 100644 --- a/Mathlib/CategoryTheory/Comma/Over.lean +++ b/Mathlib/CategoryTheory/Comma/Over.lean @@ -110,6 +110,16 @@ def isoMk {f g : Over X} (hl : f.left ≅ g.left) (hw : hl.hom ≫ g.hom = f.hom -- Porting note: simp solves this; simpNF still sees them after `-simp` (?) attribute [-simp, nolint simpNF] isoMk_hom_right_down_down isoMk_inv_right_down_down +@[reassoc (attr := simp)] +lemma hom_left_inv_left {f g : Over X} (e : f ≅ g) : + e.hom.left ≫ e.inv.left = 𝟙 f.left := by + simp [← Over.comp_left] + +@[reassoc (attr := simp)] +lemma inv_left_hom_left {f g : Over X} (e : f ≅ g) : + e.inv.left ≫ e.hom.left = 𝟙 g.left := by + simp [← Over.comp_left] + section variable (X) @@ -478,6 +488,16 @@ theorem isoMk_inv_right {f g : Under X} (hr : f.right ≅ g.right) (hw : f.hom (isoMk hr hw).inv.right = hr.inv := rfl +@[reassoc (attr := simp)] +lemma hom_right_inv_right {f g : Under X} (e : f ≅ g) : + e.hom.right ≫ e.inv.right = 𝟙 f.right := by + simp [← Under.comp_right] + +@[reassoc (attr := simp)] +lemma inv_right_hom_right {f g : Under X} (e : f ≅ g) : + e.inv.right ≫ e.hom.right = 𝟙 g.right := by + simp [← Under.comp_right] + section variable (X) diff --git a/Mathlib/RingTheory/TensorProduct/Basic.lean b/Mathlib/RingTheory/TensorProduct/Basic.lean index 58b8d3ab83f95..857c189049c66 100644 --- a/Mathlib/RingTheory/TensorProduct/Basic.lean +++ b/Mathlib/RingTheory/TensorProduct/Basic.lean @@ -902,6 +902,15 @@ theorem map_comp [Algebra S C] [IsScalarTower R S C] map (f₂.comp f₁) (g₂.comp g₁) = (map f₂ g₂).comp (map f₁ g₁) := ext (AlgHom.ext fun _ => rfl) (AlgHom.ext fun _ => rfl) +lemma map_id_comp (g₂ : E →ₐ[R] F) (g₁ : D →ₐ[R] E) : + map (AlgHom.id S A) (g₂.comp g₁) = (map (AlgHom.id S A) g₂).comp (map (AlgHom.id S A) g₁) := + ext (AlgHom.ext fun _ => rfl) (AlgHom.ext fun _ => rfl) + +lemma map_comp_id [Algebra S C] [IsScalarTower R S C] + (f₂ : B →ₐ[S] C) (f₁ : A →ₐ[S] B) : + map (f₂.comp f₁) (AlgHom.id R E) = (map f₂ (AlgHom.id R E)).comp (map f₁ (AlgHom.id R E)) := + ext (AlgHom.ext fun _ => rfl) (AlgHom.ext fun _ => rfl) + @[simp] theorem map_comp_includeLeft (f : A →ₐ[S] B) (g : C →ₐ[R] D) : (map f g).comp includeLeft = includeLeft.comp f := From 00cad59469294e95e01861e18a0026801fad3c4e Mon Sep 17 00:00:00 2001 From: grunweg Date: Tue, 26 Nov 2024 12:08:54 +0000 Subject: [PATCH 316/829] chore(100.yml): various clean-ups (#19503) Standardise how multiple authors are listed, by - removing a 'mathlib (author names)' identifier: it is apparent when results are in mathlib - always using `and` to separate multiple authors of the same formalisation (the previous state used a mix of `and` and `,`, yielding inconsistent results on the webpage). Keep using `,` to separate authors of different formalisations. --- docs/100.yaml | 28 ++++++++++++++-------------- 1 file changed, 14 insertions(+), 14 deletions(-) diff --git a/docs/100.yaml b/docs/100.yaml index 550cf7682b392..92f8c616c82ff 100644 --- a/docs/100.yaml +++ b/docs/100.yaml @@ -35,7 +35,7 @@ 9: title : The Area of a Circle decl : Theorems100.area_disc - author : James Arthur, Benjamin Davidson, and Andrew Souther + author : James Arthur and Benjamin Davidson and Andrew Souther 10: title : Euler’s Generalization of Fermat’s Little Theorem decl : Nat.ModEq.pow_totient @@ -51,7 +51,7 @@ 14: title : Euler’s Summation of 1 + (1/2)^2 + (1/3)^2 + …. decl : hasSum_zeta_two - author : Marc Masdeu, David Loeffler + author : Marc Masdeu and David Loeffler 15: title : Fundamental Theorem of Integral Calculus decls : @@ -113,7 +113,7 @@ title : Feuerbach’s Theorem 30: title : The Ballot Problem - author : Bhavik Mehta, Kexing Ying + author : Bhavik Mehta and Kexing Ying decl : Ballot.ballot_problem 31: title : Ramsey’s Theorem @@ -127,7 +127,7 @@ 34: title : Divergence of the Harmonic Series decl : Real.tendsto_sum_range_one_div_nat_succ_atTop - author : Anatole Dedecker, Yury Kudryashov + author : Anatole Dedecker and Yury Kudryashov 35: title : Taylor’s Theorem decls : @@ -157,12 +157,12 @@ 40: title : Minkowski’s Fundamental Theorem decl : MeasureTheory.exists_ne_zero_mem_lattice_of_measure_mul_two_pow_lt_measure - author : Alex J. Best, Yaël Dillies + author : Alex J. Best and Yaël Dillies 41: title : Puiseux’s Theorem 42: title : Sum of the Reciprocals of the Triangular Numbers - author : Jalex Stark, Yury Kudryashov + author : Jalex Stark and Yury Kudryashov decl : Theorems100.inverse_triangle_sum 43: title : The Isoperimetric Theorem @@ -172,7 +172,7 @@ author : Chris Hughes 45: title : The Partition Theorem - author : Bhavik Mehta, Aaron Anderson + author : Bhavik Mehta and Aaron Anderson decl : Theorems100.partition_theorem 46: title : The Solution of the General Quartic Equation @@ -310,7 +310,7 @@ decls : - sum_range_pow - sum_Ico_pow - author : mathlib (Moritz Firsching, Fabian Kruse, Ashvni Narayanan) + author : Moritz Firsching and Fabian Kruse and Ashvni Narayanan 78: title : The Cauchy-Schwarz Inequality decls : @@ -320,7 +320,7 @@ 79: title : The Intermediate Value Theorem decl : intermediate_value_Icc - author : mathlib (Rob Lewis and Chris Hughes) + author : Rob Lewis and Chris Hughes 80: title : The Fundamental Theorem of Arithmetic decls : @@ -329,14 +329,14 @@ - EuclideanDomain.to_principal_ideal_domain - UniqueFactorizationMonoid - UniqueFactorizationMonoid.factors_unique - author : mathlib (Chris Hughes) + author : Chris Hughes note : "it also has a generalized version, by showing that every Euclidean domain is a unique factorization domain, and showing that the integers form a Euclidean domain." 81: title : Divergence of the Prime Reciprocal Series decls : - Theorems100.Real.tendsto_sum_one_div_prime_atTop - not_summable_one_div_on_primes - author : Manuel Candales (archive), Michael Stoll (Mathlib) + author : Manuel Candales (archive), Michael Stoll (mathlib) 82: title : Dissection of Cubes (J.E. Littlewood’s ‘elegant’ proof) decl : Theorems100.«82».cannot_cube_a_cube @@ -344,7 +344,7 @@ 83: title : The Friendship Theorem decl : Theorems100.friendship_theorem - author : Aaron Anderson, Jalex Stark, Kyle Miller + author : Aaron Anderson and Jalex Stark and Kyle Miller 84: title : Morley’s Theorem 85: @@ -374,7 +374,7 @@ title : Stirling’s Formula decls : - Stirling.tendsto_stirlingSeq_sqrt_pi - author : mathlib (Moritz Firsching, Fabian Kruse, Nikolas Kuhn, Heather Macbeth) + author : Moritz Firsching and Fabian Kruse and Nikolas Kuhn and Heather Macbeth 91: title : The Triangle Inequality decl : norm_add_le @@ -412,7 +412,7 @@ 98: title : Bertrand’s Postulate decl : Nat.bertrand - author : Bolton Bailey, Patrick Stevens + author : Bolton Bailey and Patrick Stevens 99: title : Buffon Needle Problem links : From 7727e723796108bd9eda38dfd44c634f63ee4393 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 26 Nov 2024 12:25:44 +0000 Subject: [PATCH 317/829] chore(*): golf away from `ENNReal.toReal_le_toReal` (#19441) In most cases, use `ENNReal.toReal_mono`, since it has 1 less `ne_top` assumption. In one case, it allows to drop an assumption. --- .../Fourier/RiemannLebesgueLemma.lean | 12 ++--- Mathlib/Data/ENNReal/Real.lean | 10 ++-- .../ConditionalExpectation/CondexpL1.lean | 9 ++-- .../ConditionalExpectation/CondexpL2.lean | 3 +- .../Function/ConditionalExpectation/Real.lean | 5 +- .../Function/LpSeminorm/CompareExp.lean | 4 +- Mathlib/MeasureTheory/Function/LpSpace.lean | 10 ++-- .../Function/UniformIntegrable.lean | 8 +--- Mathlib/MeasureTheory/Integral/Bochner.lean | 7 ++- .../MeasureTheory/Integral/SetIntegral.lean | 7 ++- Mathlib/MeasureTheory/Integral/SetToL1.lean | 12 ++--- .../Measure/LevyProkhorovMetric.lean | 19 +++----- .../MeasureTheory/Measure/Portmanteau.lean | 5 +- Mathlib/Probability/BorelCantelli.lean | 6 +-- .../Kernel/Disintegration/CondCDF.lean | 6 +-- .../Kernel/Disintegration/Density.lean | 47 +++++-------------- Mathlib/Probability/Kernel/Integral.lean | 5 +- Mathlib/Probability/Moments.lean | 4 +- Mathlib/Topology/Instances/ENNReal.lean | 12 ++--- Mathlib/Topology/MetricSpace/Bounded.lean | 3 +- .../MetricSpace/HausdorffDistance.lean | 15 ++---- 21 files changed, 71 insertions(+), 138 deletions(-) diff --git a/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean b/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean index 1531edb4f4be0..08b29d45a4d25 100644 --- a/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean +++ b/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean @@ -170,20 +170,16 @@ theorem tendsto_integral_exp_inner_smul_cocompact_of_continuous_compact_support have bdA2 := norm_setIntegral_le_of_norm_le_const (hB_vol.trans_lt ENNReal.coe_lt_top) bdA ?_ swap · apply Continuous.aestronglyMeasurable - exact - continuous_norm.comp <| - Continuous.sub hf1 <| Continuous.comp hf1 <| continuous_id'.add continuous_const + exact continuous_norm.comp <| hf1.sub <| hf1.comp <| continuous_id'.add continuous_const have : ‖_‖ = ∫ v : V in A, ‖f v - f (v + i w)‖ := 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, mul_lt_mul_left hε] - rw [← ENNReal.toReal_le_toReal] at hB_vol - · refine hB_vol.trans_lt ?_ - rw [(by rfl : (↑B : ENNReal).toReal = ↑B), two_mul] - exact lt_add_of_pos_left _ hB_pos - exacts [(hB_vol.trans_lt ENNReal.coe_lt_top).ne, ENNReal.coe_lt_top.ne] + refine (ENNReal.toReal_mono ENNReal.coe_ne_top hB_vol).trans_lt ?_ + rw [ENNReal.coe_toReal, two_mul] + exact lt_add_of_pos_left _ hB_pos variable (f) diff --git a/Mathlib/Data/ENNReal/Real.lean b/Mathlib/Data/ENNReal/Real.lean index 101673394b1c5..879ef89de544e 100644 --- a/Mathlib/Data/ENNReal/Real.lean +++ b/Mathlib/Data/ENNReal/Real.lean @@ -130,13 +130,13 @@ theorem toNNReal_lt_of_lt_coe (h : a < p) : a.toNNReal < p := theorem toReal_max (hr : a ≠ ∞) (hp : b ≠ ∞) : ENNReal.toReal (max a b) = max (ENNReal.toReal a) (ENNReal.toReal b) := (le_total a b).elim - (fun h => by simp only [h, (ENNReal.toReal_le_toReal hr hp).2 h, max_eq_right]) fun h => by - simp only [h, (ENNReal.toReal_le_toReal hp hr).2 h, max_eq_left] + (fun h => by simp only [h, ENNReal.toReal_mono hp h, max_eq_right]) fun h => by + simp only [h, ENNReal.toReal_mono hr h, max_eq_left] theorem toReal_min {a b : ℝ≥0∞} (hr : a ≠ ∞) (hp : b ≠ ∞) : ENNReal.toReal (min a b) = min (ENNReal.toReal a) (ENNReal.toReal b) := - (le_total a b).elim (fun h => by simp only [h, (ENNReal.toReal_le_toReal hr hp).2 h, min_eq_left]) - fun h => by simp only [h, (ENNReal.toReal_le_toReal hp hr).2 h, min_eq_right] + (le_total a b).elim (fun h => by simp only [h, ENNReal.toReal_mono hp h, min_eq_left]) + fun h => by simp only [h, ENNReal.toReal_mono hr h, min_eq_right] theorem toReal_sup {a b : ℝ≥0∞} : a ≠ ∞ → b ≠ ∞ → (a ⊔ b).toReal = a.toReal ⊔ b.toReal := toReal_max @@ -428,7 +428,7 @@ protected theorem trichotomy₂ {p q : ℝ≥0∞} (hpq : p ≤ q) : repeat' right have hq' : 0 < q := lt_of_lt_of_le hp hpq have hp' : p < ∞ := lt_of_le_of_lt hpq hq - simp [ENNReal.toReal_le_toReal hp'.ne hq.ne, ENNReal.toReal_pos_iff, hpq, hp, hp', hq', hq] + simp [ENNReal.toReal_mono hq.ne hpq, ENNReal.toReal_pos_iff, hp, hp', hq', hq] protected theorem dichotomy (p : ℝ≥0∞) [Fact (1 ≤ p)] : p = ∞ ∨ 1 ≤ p.toReal := haveI : p = ⊤ ∨ 0 < p.toReal ∧ 1 ≤ p.toReal := by diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean index cdd2dddd698c5..b90f7b167be62 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean @@ -119,10 +119,8 @@ theorem condexpIndL1Fin_smul' [NormedSpace ℝ F] [SMulCommClass ℝ 𝕜 F] (hs theorem norm_condexpIndL1Fin_le (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : G) : ‖condexpIndL1Fin hm hs hμs x‖ ≤ (μ s).toReal * ‖x‖ := by - have : 0 ≤ ∫ a : α, ‖condexpIndL1Fin hm hs hμs x a‖ ∂μ := by positivity - rw [L1.norm_eq_integral_norm, ← ENNReal.toReal_ofReal (norm_nonneg x), ← ENNReal.toReal_mul, ← - ENNReal.toReal_ofReal this, - ENNReal.toReal_le_toReal ENNReal.ofReal_ne_top (ENNReal.mul_ne_top hμs ENNReal.ofReal_ne_top), + rw [L1.norm_eq_integral_norm, ← ENNReal.toReal_ofReal (norm_nonneg x), ← ENNReal.toReal_mul, + ← ENNReal.ofReal_le_iff_le_toReal (ENNReal.mul_ne_top hμs ENNReal.ofReal_ne_top), ofReal_integral_norm_eq_lintegral_nnnorm] swap; · rw [← memℒp_one_iff_integrable]; exact Lp.memℒp _ have h_eq : @@ -140,8 +138,7 @@ theorem condexpIndL1Fin_disjoint_union (hs : MeasurableSet s) (ht : MeasurableSe (lt_top_iff_ne_top.mpr (ENNReal.add_ne_top.mpr ⟨hμs, hμt⟩))).ne x = condexpIndL1Fin hm hs hμs x + condexpIndL1Fin hm ht hμt x := by ext1 - have hμst := - ((measure_union_le s t).trans_lt (lt_top_iff_ne_top.mpr (ENNReal.add_ne_top.mpr ⟨hμs, hμt⟩))).ne + have hμst := measure_union_ne_top hμs hμt refine (condexpIndL1Fin_ae_eq_condexpIndSMul hm (hs.union ht) hμst x).trans ?_ refine EventuallyEq.trans ?_ (Lp.coeFn_add _ _).symm have hs_eq := condexpIndL1Fin_ae_eq_condexpIndSMul hm hs hμs x diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean index 54f8441a10130..5593cd5e6e6b1 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean @@ -106,8 +106,7 @@ alias snorm_condexpL2_le := eLpNorm_condexpL2_le theorem norm_condexpL2_coe_le (hm : m ≤ m0) (f : α →₂[μ] E) : ‖(condexpL2 E 𝕜 hm f : α →₂[μ] E)‖ ≤ ‖f‖ := by rw [Lp.norm_def, Lp.norm_def, ← lpMeas_coe] - refine (ENNReal.toReal_le_toReal ?_ (Lp.eLpNorm_ne_top _)).mpr (eLpNorm_condexpL2_le hm f) - exact Lp.eLpNorm_ne_top _ + exact ENNReal.toReal_mono (Lp.eLpNorm_ne_top _) (eLpNorm_condexpL2_le hm f) theorem inner_condexpL2_left_eq_right (hm : m ≤ m0) {f g : α →₂[μ] E} : ⟪(condexpL2 E 𝕜 hm f : α →₂[μ] E), g⟫₂ = ⟪f, (condexpL2 E 𝕜 hm g : α →₂[μ] E)⟫₂ := diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean index 44d7dd73837dd..19641592ce80f 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean @@ -99,11 +99,10 @@ theorem integral_abs_condexp_le (f : α → ℝ) : ∫ x, |(μ[f|m]) x| ∂μ mul_zero] positivity rw [integral_eq_lintegral_of_nonneg_ae, integral_eq_lintegral_of_nonneg_ae] - · rw [ENNReal.toReal_le_toReal] <;> simp_rw [← Real.norm_eq_abs, ofReal_norm_eq_coe_nnnorm] + · apply ENNReal.toReal_mono <;> simp_rw [← Real.norm_eq_abs, ofReal_norm_eq_coe_nnnorm] + · exact hfint.2.ne · rw [← eLpNorm_one_eq_lintegral_nnnorm, ← eLpNorm_one_eq_lintegral_nnnorm] exact eLpNorm_one_condexp_le_eLpNorm _ - · exact integrable_condexp.2.ne - · exact hfint.2.ne · filter_upwards with x using abs_nonneg _ · simp_rw [← Real.norm_eq_abs] exact hfint.1.norm diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean index e7723bfa145de..796c3240dc8c4 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean @@ -84,7 +84,7 @@ theorem eLpNorm_le_eLpNorm_mul_rpow_measure_univ {p q : ℝ≥0∞} (hpq : p ≤ have hp_lt_top : p < ∞ := hpq.trans_lt (lt_top_iff_ne_top.mpr hq_top) have hp_pos : 0 < p.toReal := ENNReal.toReal_pos hp0_lt.ne' hp_lt_top.ne rw [eLpNorm_eq_eLpNorm' hp0_lt.ne.symm hp_lt_top.ne, eLpNorm_eq_eLpNorm' hq0_lt.ne.symm hq_top] - have hpq_real : p.toReal ≤ q.toReal := by rwa [ENNReal.toReal_le_toReal hp_lt_top.ne hq_top] + have hpq_real : p.toReal ≤ q.toReal := ENNReal.toReal_mono hq_top hpq exact eLpNorm'_le_eLpNorm'_mul_rpow_measure_univ hp_pos hpq_real hf @[deprecated (since := "2024-07-27")] @@ -155,7 +155,7 @@ theorem Memℒp.memℒp_of_exponent_le {p q : ℝ≥0∞} [IsFiniteMeasure μ] { have hp_eq_zero : p = 0 := le_antisymm (by rwa [hq_eq_zero] at hpq) (zero_le _) rw [hp_eq_zero, ENNReal.zero_toReal] at hp_pos exact (lt_irrefl _) hp_pos - have hpq_real : p.toReal ≤ q.toReal := by rwa [ENNReal.toReal_le_toReal hp_top hq_top] + have hpq_real : p.toReal ≤ q.toReal := ENNReal.toReal_mono hq_top hpq rw [eLpNorm_eq_eLpNorm' hp0 hp_top] rw [eLpNorm_eq_eLpNorm' hq0 hq_top] at hfq_lt_top exact eLpNorm'_lt_top_of_eLpNorm'_lt_top_of_exponent_le hfq_m hfq_lt_top hp_pos.le hpq_real diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index 582ddfc879f14..0d14ecf974a22 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -352,8 +352,8 @@ theorem norm_le_mul_norm_of_ae_le_mul {c : ℝ} {f : Lp E p μ} {g : Lp F p μ} theorem norm_le_norm_of_ae_le {f : Lp E p μ} {g : Lp F p μ} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ ‖g x‖) : ‖f‖ ≤ ‖g‖ := by - rw [norm_def, norm_def, ENNReal.toReal_le_toReal (eLpNorm_ne_top _) (eLpNorm_ne_top _)] - exact eLpNorm_mono_ae h + rw [norm_def, norm_def] + exact ENNReal.toReal_mono (eLpNorm_ne_top _) (eLpNorm_mono_ae h) theorem mem_Lp_of_nnnorm_ae_le_mul {c : ℝ≥0} {f : α →ₘ[μ] E} {g : Lp F p μ} (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ c * ‖g x‖₊) : f ∈ Lp E p μ := @@ -1619,9 +1619,9 @@ theorem ae_tendsto_of_cauchy_eLpNorm [CompleteSpace E] {f : ℕ → α → E} refine cauchySeq_of_le_tendsto_0 (fun n => (B n).toReal) ?_ ?_ · intro n m N hnN hmN specialize hx N n m hnN hmN - rw [_root_.dist_eq_norm, ← ENNReal.toReal_ofReal (norm_nonneg _), - ENNReal.toReal_le_toReal ENNReal.ofReal_ne_top (ENNReal.ne_top_of_tsum_ne_top hB N)] - rw [← ofReal_norm_eq_coe_nnnorm] at hx + rw [_root_.dist_eq_norm, + ← ENNReal.ofReal_le_iff_le_toReal (ENNReal.ne_top_of_tsum_ne_top hB N), + ofReal_norm_eq_coe_nnnorm] exact hx.le · rw [← ENNReal.zero_toReal] exact diff --git a/Mathlib/MeasureTheory/Function/UniformIntegrable.lean b/Mathlib/MeasureTheory/Function/UniformIntegrable.lean index 1f933577f78c9..c7c8a68fdf2d0 100644 --- a/Mathlib/MeasureTheory/Function/UniformIntegrable.lean +++ b/Mathlib/MeasureTheory/Function/UniformIntegrable.lean @@ -517,12 +517,8 @@ theorem tendsto_Lp_finite_of_tendsto_ae_of_meas [IsFiniteMeasure μ] (hp : 1 ≤ exact ⟨0, fun n _ => by simp [h]⟩ by_cases hμ : μ = 0 · exact ⟨0, fun n _ => by simp [hμ]⟩ - have hε' : 0 < ε.toReal / 3 := - div_pos (ENNReal.toReal_pos (gt_iff_lt.1 hε).ne.symm h.ne) (by norm_num) - have hdivp : 0 ≤ 1 / p.toReal := by - refine one_div_nonneg.2 ?_ - rw [← ENNReal.zero_toReal, ENNReal.toReal_le_toReal ENNReal.zero_ne_top hp'] - exact le_trans (zero_le _) hp + have hε' : 0 < ε.toReal / 3 := div_pos (ENNReal.toReal_pos hε.ne' h.ne) (by norm_num) + have hdivp : 0 ≤ 1 / p.toReal := by positivity have hpow : 0 < measureUnivNNReal μ ^ (1 / p.toReal) := Real.rpow_pos_of_pos (measureUnivNNReal_pos hμ) _ obtain ⟨δ₁, hδ₁, heLpNorm₁⟩ := hui hε' diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index 2cf72496737ad..a3a28a6bca1d4 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -1352,10 +1352,9 @@ theorem integral_mono_measure {f : α → ℝ} {ν} (hle : μ ≤ ν) (hf : 0 (hfi : Integrable f ν) : ∫ a, f a ∂μ ≤ ∫ a, f a ∂ν := by have hfi' : Integrable f μ := hfi.mono_measure hle have hf' : 0 ≤ᵐ[μ] f := hle.absolutelyContinuous hf - rw [integral_eq_lintegral_of_nonneg_ae hf' hfi'.1, integral_eq_lintegral_of_nonneg_ae hf hfi.1, - ENNReal.toReal_le_toReal] - exacts [lintegral_mono' hle le_rfl, ((hasFiniteIntegral_iff_ofReal hf').1 hfi'.2).ne, - ((hasFiniteIntegral_iff_ofReal hf).1 hfi.2).ne] + rw [integral_eq_lintegral_of_nonneg_ae hf' hfi'.1, integral_eq_lintegral_of_nonneg_ae hf hfi.1] + refine ENNReal.toReal_mono ?_ (lintegral_mono' hle le_rfl) + exact ((hasFiniteIntegral_iff_ofReal hf).1 hfi.2).ne theorem norm_integral_le_integral_norm (f : α → G) : ‖∫ a, f a ∂μ‖ ≤ ∫ a, ‖f a‖ ∂μ := by have le_ae : ∀ᵐ a ∂μ, 0 ≤ ‖f a‖ := Eventually.of_forall fun a => norm_nonneg _ diff --git a/Mathlib/MeasureTheory/Integral/SetIntegral.lean b/Mathlib/MeasureTheory/Integral/SetIntegral.lean index 24dbbbb85ffe7..4636afb0369c5 100644 --- a/Mathlib/MeasureTheory/Integral/SetIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/SetIntegral.lean @@ -1019,10 +1019,9 @@ theorem Lp_toLp_restrict_smul (c : 𝕜) (f : Lp F p μ) (s : Set X) : `(Lp.memℒp f).restrict s).toLp f`. This map is non-expansive. -/ theorem norm_Lp_toLp_restrict_le (s : Set X) (f : Lp E p μ) : ‖((Lp.memℒp f).restrict s).toLp f‖ ≤ ‖f‖ := by - rw [Lp.norm_def, Lp.norm_def, ENNReal.toReal_le_toReal (Lp.eLpNorm_ne_top _) - (Lp.eLpNorm_ne_top _)] - apply (le_of_eq _).trans (eLpNorm_mono_measure _ (Measure.restrict_le_self (s := s))) - exact eLpNorm_congr_ae (Memℒp.coeFn_toLp _) + rw [Lp.norm_def, Lp.norm_def, eLpNorm_congr_ae (Memℒp.coeFn_toLp _)] + refine ENNReal.toReal_mono (Lp.eLpNorm_ne_top _) ?_ + exact eLpNorm_mono_measure _ Measure.restrict_le_self variable (X F 𝕜) in /-- Continuous linear map sending a function of `Lp F p μ` to the same function in diff --git a/Mathlib/MeasureTheory/Integral/SetToL1.lean b/Mathlib/MeasureTheory/Integral/SetToL1.lean index 456069af75266..2a0d2d70bd546 100644 --- a/Mathlib/MeasureTheory/Integral/SetToL1.lean +++ b/Mathlib/MeasureTheory/Integral/SetToL1.lean @@ -1422,20 +1422,14 @@ theorem continuous_L1_toL1 {μ' : Measure α} (c' : ℝ≥0∞) (hc' : c' ≠ rw [this] have h_eLpNorm_ne_top : eLpNorm (⇑g - ⇑f) 1 μ ≠ ∞ := by rw [← eLpNorm_congr_ae (Lp.coeFn_sub _ _)]; exact Lp.eLpNorm_ne_top _ - have h_eLpNorm_ne_top' : eLpNorm (⇑g - ⇑f) 1 μ' ≠ ∞ := by - refine ((eLpNorm_mono_measure _ hμ'_le).trans_lt ?_).ne - rw [eLpNorm_smul_measure_of_ne_zero hc'0, smul_eq_mul] - refine ENNReal.mul_lt_top ?_ h_eLpNorm_ne_top.lt_top - simp [hc'.lt_top, hc'0] calc (eLpNorm (⇑g - ⇑f) 1 μ').toReal ≤ (c' * eLpNorm (⇑g - ⇑f) 1 μ).toReal := by - rw [toReal_le_toReal h_eLpNorm_ne_top' (ENNReal.mul_ne_top hc' h_eLpNorm_ne_top)] - refine (eLpNorm_mono_measure (⇑g - ⇑f) hμ'_le).trans ?_ + refine toReal_mono (ENNReal.mul_ne_top hc' h_eLpNorm_ne_top) ?_ + refine (eLpNorm_mono_measure (⇑g - ⇑f) hμ'_le).trans_eq ?_ rw [eLpNorm_smul_measure_of_ne_zero hc'0, smul_eq_mul] simp _ = c'.toReal * (eLpNorm (⇑g - ⇑f) 1 μ).toReal := toReal_mul - _ ≤ c'.toReal * (ε / 2 / c'.toReal) := - (mul_le_mul le_rfl hfg.le toReal_nonneg toReal_nonneg) + _ ≤ c'.toReal * (ε / 2 / c'.toReal) := by gcongr _ = ε / 2 := by refine mul_div_cancel₀ (ε / 2) ?_; rw [Ne, toReal_eq_zero_iff]; simp [hc', hc'0] _ < ε := half_lt_self hε_pos diff --git a/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean b/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean index 02ed4b6805a38..3fb7d9b953875 100644 --- a/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean +++ b/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean @@ -180,14 +180,10 @@ lemma levyProkhorovDist_comm (μ ν : Measure Ω) : lemma levyProkhorovDist_triangle [OpensMeasurableSpace Ω] (μ ν κ : Measure Ω) [IsFiniteMeasure μ] [IsFiniteMeasure ν] [IsFiniteMeasure κ] : levyProkhorovDist μ κ ≤ levyProkhorovDist μ ν + levyProkhorovDist ν κ := by - have dμκ_finite := (levyProkhorovEDist_lt_top μ κ).ne have dμν_finite := (levyProkhorovEDist_lt_top μ ν).ne have dνκ_finite := (levyProkhorovEDist_lt_top ν κ).ne - convert (ENNReal.toReal_le_toReal (a := levyProkhorovEDist μ κ) - (b := levyProkhorovEDist μ ν + levyProkhorovEDist ν κ) - _ _).mpr <| levyProkhorovEDist_triangle μ ν κ + convert ENNReal.toReal_mono ?_ <| levyProkhorovEDist_triangle μ ν κ · simp only [levyProkhorovDist, ENNReal.toReal_add dμν_finite dνκ_finite] - · exact dμκ_finite · exact ENNReal.add_ne_top.mpr ⟨dμν_finite, dνκ_finite⟩ /-- A type synonym, to be used for `Measure α`, `FiniteMeasure α`, or `ProbabilityMeasure α`, @@ -366,9 +362,8 @@ lemma BoundedContinuousFunction.integral_le_of_levyProkhorovEDist_lt (μ ν : Me have key : (fun (t : ℝ) ↦ ENNReal.toReal (μ {a | t ≤ f a})) ≤ (fun (t : ℝ) ↦ ENNReal.toReal (ν (thickening ε {a | t ≤ f a})) + ε) := by intro t - convert (ENNReal.toReal_le_toReal (measure_ne_top _ _) ?_).mpr - <| left_measure_le_of_levyProkhorovEDist_lt hμν (B := {a | t ≤ f a}) - (f.continuous.measurable measurableSet_Ici) + convert ENNReal.toReal_mono ?_ <| left_measure_le_of_levyProkhorovEDist_lt hμν + (B := {a | t ≤ f a}) (f.continuous.measurable measurableSet_Ici) · rw [ENNReal.toReal_add (measure_ne_top ν _) ofReal_ne_top, ENNReal.toReal_ofReal ε_pos.le] · exact ENNReal.add_ne_top.mpr ⟨measure_ne_top ν _, ofReal_ne_top⟩ have intble₁ : IntegrableOn (fun t ↦ ENNReal.toReal (μ {a | t ≤ f a})) (Ioc 0 ‖f‖) := by @@ -377,8 +372,7 @@ lemma BoundedContinuousFunction.integral_le_of_levyProkhorovEDist_lt (μ ν : Me exact fun _ _ hst ↦ measure_mono (fun _ h ↦ hst.trans h) · apply Eventually.of_forall <| fun t ↦ ?_ simp only [Real.norm_eq_abs, abs_toReal] - exact (ENNReal.toReal_le_toReal (measure_ne_top _ _) (measure_ne_top _ _)).mpr - <| measure_mono (subset_univ _) + exact ENNReal.toReal_mono (measure_ne_top _ _) <| measure_mono (subset_univ _) have intble₂ : IntegrableOn (fun t ↦ ENNReal.toReal (ν (thickening ε {a | t ≤ f a}))) (Ioc 0 ‖f‖) := by apply Measure.integrableOn_of_bounded (M := ENNReal.toReal (ν univ)) measure_Ioc_lt_top.ne @@ -386,8 +380,7 @@ lemma BoundedContinuousFunction.integral_le_of_levyProkhorovEDist_lt (μ ν : Me exact fun _ _ hst ↦ measure_mono <| thickening_subset_of_subset ε (fun _ h ↦ hst.trans h) · apply Eventually.of_forall <| fun t ↦ ?_ simp only [Real.norm_eq_abs, abs_toReal] - exact (ENNReal.toReal_le_toReal (measure_ne_top _ _) (measure_ne_top _ _)).mpr - <| measure_mono (subset_univ _) + exact ENNReal.toReal_mono (measure_ne_top _ _) <| measure_mono (subset_univ _) apply le_trans (setIntegral_mono (s := Ioc 0 ‖f‖) ?_ ?_ key) · rw [integral_add] · apply add_le_add_left @@ -415,7 +408,7 @@ lemma tendsto_integral_meas_thickening_le (f : Ω →ᵇ ℝ) · apply Eventually.of_forall (fun i ↦ ?_) apply Eventually.of_forall (fun t ↦ ?_) simp only [Real.norm_eq_abs, NNReal.abs_eq, Pi.one_apply] - exact (ENNReal.toReal_le_toReal (measure_ne_top _ _) one_ne_top).mpr prob_le_one + exact ENNReal.toReal_mono one_ne_top prob_le_one · have aux : IsFiniteMeasure (volume.restrict A) := ⟨by simp [lt_top_iff_ne_top, A_finmeas]⟩ apply integrable_const · apply Eventually.of_forall (fun t ↦ ?_) diff --git a/Mathlib/MeasureTheory/Measure/Portmanteau.lean b/Mathlib/MeasureTheory/Measure/Portmanteau.lean index 72fe3305828be..6cb6af9e29fd8 100644 --- a/Mathlib/MeasureTheory/Measure/Portmanteau.lean +++ b/Mathlib/MeasureTheory/Measure/Portmanteau.lean @@ -481,7 +481,7 @@ lemma lintegral_le_liminf_lintegral_of_forall_isOpen_measure_le_liminf_measure measure_mono (fun ω hω ↦ lt_of_le_of_lt hst hω))) lemma integral_le_liminf_integral_of_forall_isOpen_measure_le_liminf_measure - {μ : Measure Ω} [IsProbabilityMeasure μ] {μs : ℕ → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] + {μ : Measure Ω} {μs : ℕ → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] {f : Ω →ᵇ ℝ} (f_nn : 0 ≤ f) (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : ∫ x, (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i)) := by @@ -489,7 +489,7 @@ lemma integral_le_liminf_integral_of_forall_isOpen_measure_le_liminf_measure f.continuous f_nn h_opens rw [@integral_eq_lintegral_of_nonneg_ae Ω _ μ f (Eventually.of_forall f_nn) f.continuous.measurable.aestronglyMeasurable] - convert (ENNReal.toReal_le_toReal ?_ ?_).mpr same + convert ENNReal.toReal_mono ?_ same · simp only [fun i ↦ @integral_eq_lintegral_of_nonneg_ae Ω _ (μs i) f (Eventually.of_forall f_nn) f.continuous.measurable.aestronglyMeasurable] let g := BoundedContinuousFunction.comp _ Real.lipschitzWith_toNNReal f @@ -497,7 +497,6 @@ lemma integral_le_liminf_integral_of_forall_isOpen_measure_le_liminf_measure simpa only [coe_nnreal_ennreal_nndist, measure_univ, mul_one, ge_iff_le] using BoundedContinuousFunction.lintegral_le_edist_mul (μ := μs i) g apply ENNReal.liminf_toReal_eq ENNReal.coe_ne_top (Eventually.of_forall bound) - · exact (f.lintegral_of_real_lt_top μ).ne · apply ne_of_lt have obs := fun (i : ℕ) ↦ @BoundedContinuousFunction.lintegral_nnnorm_le Ω _ _ (μs i) ℝ _ f simp only [measure_univ, mul_one] at obs diff --git a/Mathlib/Probability/BorelCantelli.lean b/Mathlib/Probability/BorelCantelli.lean index be4a2299f904e..c94f855abd159 100644 --- a/Mathlib/Probability/BorelCantelli.lean +++ b/Mathlib/Probability/BorelCantelli.lean @@ -92,9 +92,9 @@ theorem measure_limsup_eq_one {s : ℕ → Set Ω} (hsm : ∀ n, MeasurableSet ( · refine ⟨n, ?_⟩ rw [ENNReal.toReal_sum] exact fun _ _ => measure_ne_top _ _ - · rw [not_lt, ← ENNReal.toReal_le_toReal (ENNReal.sum_ne_top.2 _) ENNReal.coe_ne_top] - · exact hB.trans (by simp) - · exact fun _ _ => measure_ne_top _ _ + · rwa [not_lt, ENNReal.ofNNReal_toNNReal, ENNReal.le_ofReal_iff_toReal_le] + · simp + · exact le_trans (by positivity) hB end BorelCantelli diff --git a/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean b/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean index b52e446ac654c..9d6ed97e8708e 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean @@ -208,10 +208,8 @@ lemma isRatCondKernelCDFAux_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure (Kernel.const Unit ρ) (Kernel.const Unit ρ.fst) where measurable := measurable_preCDF'.comp measurable_snd mono' a r r' hrr' := by - filter_upwards [monotone_preCDF ρ, preCDF_le_one ρ] with a h1 h2 - have h_ne_top : ∀ r, preCDF ρ r a ≠ ∞ := fun r ↦ ((h2 r).trans_lt ENNReal.one_lt_top).ne - rw [ENNReal.toReal_le_toReal (h_ne_top _) (h_ne_top _)] - exact h1 hrr' + filter_upwards [monotone_preCDF ρ, preCDF_le_one ρ] with a h₁ h₂ + exact ENNReal.toReal_mono ((h₂ _).trans_lt ENNReal.one_lt_top).ne (h₁ hrr') nonneg' _ q := by simp le_one' a q := by simp only [Kernel.const_apply, forall_const] diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index a3451dffcbecf..94af1027b716b 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -309,18 +309,11 @@ lemma densityProcess_mono_set (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (x : γ {s s' : Set β} (h : s ⊆ s') : densityProcess κ ν n a x s ≤ densityProcess κ ν n a x s' := by unfold densityProcess - by_cases h0 : ν a (countablePartitionSet n x) = 0 - · rw [h0, ENNReal.toReal_div, ENNReal.toReal_div] - simp - have h_ne_top : ∀ s, - κ a (countablePartitionSet n x ×ˢ s) / ν a (countablePartitionSet n x) ≠ ⊤ := by - intro s - rw [ne_eq, ENNReal.div_eq_top] - simp only [ne_eq, h0, and_false, false_or, not_and, not_not] - refine fun h_top ↦ eq_top_mono ?_ h_top - exact meas_countablePartitionSet_le_of_fst_le hκν n a x s - rw [ENNReal.toReal_le_toReal (h_ne_top s) (h_ne_top s')] - gcongr + obtain h₀ | h₀ := eq_or_ne (ν a (countablePartitionSet n x)) 0 + · simp [h₀] + · gcongr + simp only [ne_eq, ENNReal.div_eq_top, h₀, and_false, false_or, not_and, not_not] + exact eq_top_mono (meas_countablePartitionSet_le_of_fst_le hκν n a x s') lemma densityProcess_mono_kernel_left {κ' : Kernel α (γ × β)} (hκκ' : κ ≤ κ') (hκ'ν : fst κ' ≤ ν) (n : ℕ) (a : α) (x : γ) (s : Set β) : @@ -331,16 +324,10 @@ lemma densityProcess_mono_kernel_left {κ' : Kernel α (γ × β)} (hκκ' : κ simp have h_le : κ' a (countablePartitionSet n x ×ˢ s) ≤ ν a (countablePartitionSet n x) := meas_countablePartitionSet_le_of_fst_le hκ'ν n a x s - rw [ENNReal.toReal_le_toReal] - · gcongr - exact hκκ' _ _ - · rw [ne_eq, ENNReal.div_eq_top] - simp only [ne_eq, h0, and_false, false_or, not_and, not_not] - refine fun h_top ↦ eq_top_mono ?_ h_top - exact (hκκ' _ _).trans h_le - · rw [ne_eq, ENNReal.div_eq_top] - simp only [ne_eq, h0, and_false, false_or, not_and, not_not] + gcongr + · simp only [ne_eq, ENNReal.div_eq_top, h0, and_false, false_or, not_and, not_not] exact fun h_top ↦ eq_top_mono h_le h_top + · apply hκκ' lemma densityProcess_antitone_kernel_right {ν' : Kernel α γ} (hνν' : ν ≤ ν') (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (x : γ) (s : Set β) : @@ -350,16 +337,10 @@ lemma densityProcess_antitone_kernel_right {ν' : Kernel α γ} meas_countablePartitionSet_le_of_fst_le hκν n a x s by_cases h0 : ν a (countablePartitionSet n x) = 0 · simp [le_antisymm (h_le.trans h0.le) zero_le', h0] - have h0' : ν' a (countablePartitionSet n x) ≠ 0 := - fun h ↦ h0 (le_antisymm ((hνν' _ _).trans h.le) zero_le') - rw [ENNReal.toReal_le_toReal] - · gcongr - exact hνν' _ _ - · simp only [ne_eq, ENNReal.div_eq_top, h0', and_false, false_or, not_and, not_not] - refine fun h_top ↦ eq_top_mono ?_ h_top - exact h_le.trans (hνν' _ _) + gcongr · simp only [ne_eq, ENNReal.div_eq_top, h0, and_false, false_or, not_and, not_not] exact fun h_top ↦ eq_top_mono h_le h_top + · apply hνν' @[simp] lemma densityProcess_empty (κ : Kernel α (γ × β)) (ν : Kernel α γ) (n : ℕ) (a : α) (x : γ) : @@ -384,10 +365,7 @@ lemma tendsto_densityProcess_atTop_empty_of_antitone (κ : Kernel α (γ × β)) have : Tendsto (fun m ↦ κ a (countablePartitionSet n x ×ˢ seq m)) atTop (𝓝 ((κ a) (⋂ n_1, countablePartitionSet n x ×ˢ seq n_1))) := by apply tendsto_measure_iInter_atTop - · intro - -- TODO: why doesn't `measurability` work without this hint? - apply MeasurableSet.nullMeasurableSet - measurability + · measurability · exact fun _ _ h ↦ prod_mono_right <| hseq h · exact ⟨0, measure_ne_top _ _⟩ simpa only [← prod_iInter, hseq_iInter] using this @@ -603,8 +581,7 @@ lemma setIntegral_density (hκν : fst κ ≤ ν) [IsFiniteKernel ν] ENNReal.toReal_sub_of_le (measure_mono (by intro x; simp)) (measure_ne_top _ _)] rw [eq_tsub_iff_add_eq_of_le, add_comm] · exact h - · rw [ENNReal.toReal_le_toReal (measure_ne_top _ _) (measure_ne_top _ _)] - exact measure_mono (by intro x; simp) + · gcongr <;> simp · intro f hf_disj hf h_eq rw [integral_iUnion _ hf_disj (integrable_density hκν _ hs).integrableOn] · simp_rw [h_eq] diff --git a/Mathlib/Probability/Kernel/Integral.lean b/Mathlib/Probability/Kernel/Integral.lean index e97ef25407719..d586ae6a66158 100644 --- a/Mathlib/Probability/Kernel/Integral.lean +++ b/Mathlib/Probability/Kernel/Integral.lean @@ -26,9 +26,8 @@ lemma IsFiniteKernel.integrable (μ : Measure α) [IsFiniteMeasure μ] refine Integrable.mono' (integrable_const (IsFiniteKernel.bound κ).toReal) ((κ.measurable_coe hs).ennreal_toReal.aestronglyMeasurable) (ae_of_all μ fun x ↦ ?_) - rw [Real.norm_eq_abs, abs_of_nonneg ENNReal.toReal_nonneg, - ENNReal.toReal_le_toReal (measure_ne_top _ _) (IsFiniteKernel.bound_ne_top _)] - exact Kernel.measure_le_bound _ _ _ + rw [Real.norm_eq_abs, abs_of_nonneg ENNReal.toReal_nonneg] + exact ENNReal.toReal_mono (IsFiniteKernel.bound_ne_top _) (Kernel.measure_le_bound _ _ _) lemma IsMarkovKernel.integrable (μ : Measure α) [IsFiniteMeasure μ] (κ : Kernel α β) [IsMarkovKernel κ] {s : Set β} (hs : MeasurableSet s) : diff --git a/Mathlib/Probability/Moments.lean b/Mathlib/Probability/Moments.lean index 1433dc8119e43..5cd7261ecfd42 100644 --- a/Mathlib/Probability/Moments.lean +++ b/Mathlib/Probability/Moments.lean @@ -289,8 +289,8 @@ theorem measure_ge_le_exp_mul_mgf [IsFiniteMeasure μ] (ε : ℝ) (ht : 0 ≤ t) rcases ht.eq_or_lt with ht_zero_eq | ht_pos · rw [ht_zero_eq.symm] simp only [neg_zero, zero_mul, exp_zero, mgf_zero', one_mul] - rw [ENNReal.toReal_le_toReal (measure_ne_top μ _) (measure_ne_top μ _)] - exact measure_mono (Set.subset_univ _) + gcongr + exacts [measure_ne_top _ _, Set.subset_univ _] calc (μ {ω | ε ≤ X ω}).toReal = (μ {ω | exp (t * ε) ≤ exp (t * X ω)}).toReal := by congr with ω diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index 029ada7f6a78a..418685fd8c5d4 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -1318,19 +1318,17 @@ lemma truncateToReal_eq_toReal {t x : ℝ≥0∞} (t_ne_top : t ≠ ∞) (x_le : lemma truncateToReal_le {t : ℝ≥0∞} (t_ne_top : t ≠ ∞) {x : ℝ≥0∞} : truncateToReal t x ≤ t.toReal := by rw [truncateToReal] - apply (toReal_le_toReal _ t_ne_top).mpr (min_le_left t x) - simp_all only [ne_eq, min_eq_top, false_and, not_false_eq_true] + gcongr + exacts [t_ne_top, min_le_left t x] lemma truncateToReal_nonneg {t x : ℝ≥0∞} : 0 ≤ truncateToReal t x := toReal_nonneg /-- The truncated cast `ENNReal.truncateToReal t : ℝ≥0∞ → ℝ` is monotone when `t ≠ ∞`. -/ lemma monotone_truncateToReal {t : ℝ≥0∞} (t_ne_top : t ≠ ∞) : Monotone (truncateToReal t) := by intro x y x_le_y - have obs_x : min t x ≠ ∞ := by - simp_all only [ne_eq, min_eq_top, false_and, not_false_eq_true] - have obs_y : min t y ≠ ∞ := by - simp_all only [ne_eq, min_eq_top, false_and, not_false_eq_true] - exact (ENNReal.toReal_le_toReal obs_x obs_y).mpr (min_le_min_left t x_le_y) + simp only [truncateToReal] + gcongr + exact ne_top_of_le_ne_top t_ne_top (min_le_left _ _) /-- The truncated cast `ENNReal.truncateToReal t : ℝ≥0∞ → ℝ` is continuous when `t ≠ ∞`. -/ lemma continuous_truncateToReal {t : ℝ≥0∞} (t_ne_top : t ≠ ∞) : Continuous (truncateToReal t) := by diff --git a/Mathlib/Topology/MetricSpace/Bounded.lean b/Mathlib/Topology/MetricSpace/Bounded.lean index ebe5246e8b3d2..ab85f93d8dea6 100644 --- a/Mathlib/Topology/MetricSpace/Bounded.lean +++ b/Mathlib/Topology/MetricSpace/Bounded.lean @@ -392,8 +392,7 @@ theorem diam_le_of_forall_dist_le_of_nonempty (hs : s.Nonempty) {C : ℝ} theorem dist_le_diam_of_mem' (h : EMetric.diam s ≠ ⊤) (hx : x ∈ s) (hy : y ∈ s) : dist x y ≤ diam s := by rw [diam, dist_edist] - rw [ENNReal.toReal_le_toReal (edist_ne_top _ _) h] - exact EMetric.edist_le_diam_of_mem hx hy + exact ENNReal.toReal_mono h <| EMetric.edist_le_diam_of_mem hx hy /-- Characterize the boundedness of a set in terms of the finiteness of its emetric.diameter. -/ theorem isBounded_iff_ediam_ne_top : IsBounded s ↔ EMetric.diam s ≠ ⊤ := diff --git a/Mathlib/Topology/MetricSpace/HausdorffDistance.lean b/Mathlib/Topology/MetricSpace/HausdorffDistance.lean index 698c804dacc7e..cb1ae246d0e34 100644 --- a/Mathlib/Topology/MetricSpace/HausdorffDistance.lean +++ b/Mathlib/Topology/MetricSpace/HausdorffDistance.lean @@ -665,24 +665,15 @@ theorem hausdorffDist_empty' : hausdorffDist ∅ s = 0 := by simp [hausdorffDist in each set to the other set -/ theorem hausdorffDist_le_of_infDist {r : ℝ} (hr : 0 ≤ r) (H1 : ∀ x ∈ s, infDist x t ≤ r) (H2 : ∀ x ∈ t, infDist x s ≤ r) : hausdorffDist s t ≤ r := by - by_cases h1 : hausdorffEdist s t = ⊤ - · rwa [hausdorffDist, h1, ENNReal.top_toReal] rcases s.eq_empty_or_nonempty with hs | hs · rwa [hs, hausdorffDist_empty'] rcases t.eq_empty_or_nonempty with ht | ht · rwa [ht, hausdorffDist_empty] have : hausdorffEdist s t ≤ ENNReal.ofReal r := by apply hausdorffEdist_le_of_infEdist _ _ - · intro x hx - have I := H1 x hx - rwa [infDist, ← ENNReal.toReal_ofReal hr, - ENNReal.toReal_le_toReal (infEdist_ne_top ht) ENNReal.ofReal_ne_top] at I - · intro x hx - have I := H2 x hx - rwa [infDist, ← ENNReal.toReal_ofReal hr, - ENNReal.toReal_le_toReal (infEdist_ne_top hs) ENNReal.ofReal_ne_top] at I - rwa [hausdorffDist, ← ENNReal.toReal_ofReal hr, - ENNReal.toReal_le_toReal h1 ENNReal.ofReal_ne_top] + · simpa only [infDist, ← ENNReal.le_ofReal_iff_toReal_le (infEdist_ne_top ht) hr] using H1 + · simpa only [infDist, ← ENNReal.le_ofReal_iff_toReal_le (infEdist_ne_top hs) hr] using H2 + exact ENNReal.toReal_le_of_le_ofReal hr this /-- Bounding the Hausdorff distance by exhibiting, for any point in each set, another point in the other set at controlled distance -/ From 5d0ac0445585738ce54244d28a68e9abd4efffa5 Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Tue, 26 Nov 2024 12:47:44 +0000 Subject: [PATCH 318/829] feat(CategoryTheory): split up a product into a binary product (#19199) --- Mathlib.lean | 1 + .../CategoryTheory/Limits/Shapes/PiProd.lean | 69 +++++++++++++++++++ 2 files changed, 70 insertions(+) create mode 100644 Mathlib/CategoryTheory/Limits/Shapes/PiProd.lean diff --git a/Mathlib.lean b/Mathlib.lean index fb9f9c436a030..3741c55a437ae 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1862,6 +1862,7 @@ import Mathlib.CategoryTheory.Limits.Shapes.Kernels import Mathlib.CategoryTheory.Limits.Shapes.Multiequalizer import Mathlib.CategoryTheory.Limits.Shapes.NormalMono.Basic import Mathlib.CategoryTheory.Limits.Shapes.NormalMono.Equalizers +import Mathlib.CategoryTheory.Limits.Shapes.PiProd import Mathlib.CategoryTheory.Limits.Shapes.Products import Mathlib.CategoryTheory.Limits.Shapes.Pullback.Assoc import Mathlib.CategoryTheory.Limits.Shapes.Pullback.CommSq diff --git a/Mathlib/CategoryTheory/Limits/Shapes/PiProd.lean b/Mathlib/CategoryTheory/Limits/Shapes/PiProd.lean new file mode 100644 index 0000000000000..6027232e5928d --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Shapes/PiProd.lean @@ -0,0 +1,69 @@ +/- +Copyright (c) 2024 Dagur Asgeirsson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Dagur Asgeirsson +-/ +import Mathlib.CategoryTheory.Limits.Shapes.BinaryProducts +import Mathlib.CategoryTheory.Limits.Shapes.Products +/-! + +# A product as a binary product + +We write a product indexed by `I` as a binary product of the products indexed by a subset of `I` +and its complement. + +-/ + +namespace CategoryTheory.Limits + +variable {C I : Type*} [Category C] {X Y : I → C} + (f : (i : I) → X i ⟶ Y i) (P : I → Prop) [∀ i, Decidable (P i)] + [HasProduct X] [HasProduct Y] + [HasProduct (fun (i : {x : I // P x}) ↦ X i.val)] + [HasProduct (fun (i : {x : I // ¬ P x}) ↦ X i.val)] + [HasProduct (fun (i : {x : I // P x}) ↦ Y i.val)] + [HasProduct (fun (i : {x : I // ¬ P x}) ↦ Y i.val)] + +variable (X) in +/-- +The projection maps of a product to the products indexed by a subset and its complement, as a +binary fan. +-/ +noncomputable def Pi.binaryFanOfProp : BinaryFan (∏ᶜ (fun (i : {x : I // P x}) ↦ X i.val)) + (∏ᶜ (fun (i : {x : I // ¬ P x}) ↦ X i.val)) := + BinaryFan.mk (P := ∏ᶜ X) (Pi.map' Subtype.val fun _ ↦ 𝟙 _) + (Pi.map' Subtype.val fun _ ↦ 𝟙 _) + +variable (X) in +/-- +A product indexed by `I` is a binary product of the products indexed by a subset of `I` and its +complement. +-/ +noncomputable def Pi.binaryFanOfPropIsLimit : IsLimit (Pi.binaryFanOfProp X P) := + BinaryFan.isLimitMk + (fun s ↦ Pi.lift fun b ↦ if h : P b then + s.π.app ⟨WalkingPair.left⟩ ≫ Pi.π (fun (i : {x : I // P x}) ↦ X i.val) ⟨b, h⟩ else + s.π.app ⟨WalkingPair.right⟩ ≫ Pi.π (fun (i : {x : I // ¬ P x}) ↦ X i.val) ⟨b, h⟩) + (by aesop) (by aesop) + (fun _ _ h₁ h₂ ↦ Pi.hom_ext _ _ fun b ↦ by + by_cases h : P b + · simp [← h₁, dif_pos h] + · simp [← h₂, dif_neg h]) + +lemma hasBinaryProduct_of_products : HasBinaryProduct (∏ᶜ (fun (i : {x : I // P x}) ↦ X i.val)) + (∏ᶜ (fun (i : {x : I // ¬ P x}) ↦ X i.val)) := + ⟨Pi.binaryFanOfProp X P, Pi.binaryFanOfPropIsLimit X P⟩ + +attribute [local instance] hasBinaryProduct_of_products + +lemma Pi.map_eq_prod_map : Pi.map f = + ((Pi.binaryFanOfPropIsLimit X P).conePointUniqueUpToIso (prodIsProd _ _)).hom ≫ + prod.map (Pi.map (fun (i : {x : I // P x}) ↦ f i.val)) + (Pi.map (fun (i : {x : I // ¬ P x}) ↦ f i.val)) ≫ + ((Pi.binaryFanOfPropIsLimit Y P).conePointUniqueUpToIso (prodIsProd _ _)).inv := by + rw [← Category.assoc, Iso.eq_comp_inv] + dsimp only [IsLimit.conePointUniqueUpToIso, binaryFanOfProp, prodIsProd] + apply prod.hom_ext + all_goals aesop_cat + +end CategoryTheory.Limits From b4b0314cc962a0b9d4cac132e7a1f742d7e640f3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 26 Nov 2024 12:47:45 +0000 Subject: [PATCH 319/829] feat: more `eLpNorm` API (#19293) Make more arguments to rewriting lemmas explicit, more arguments to the other lemmas implicit. Add various lemmas. From LeanAPAP --- .../Function/LpSeminorm/Basic.lean | 96 ++++++++++++++----- Mathlib/MeasureTheory/Function/LpSpace.lean | 4 +- 2 files changed, 73 insertions(+), 27 deletions(-) diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean index 71c2449fa8236..4c9f13b1bf99b 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne, Sébastien Gouëzel -/ import Mathlib.Analysis.NormedSpace.IndicatorFunction +import Mathlib.Data.Fintype.Order import Mathlib.MeasureTheory.Function.EssSup import Mathlib.MeasureTheory.Function.AEEqFun import Mathlib.MeasureTheory.Function.SpecialFunctions.Basic @@ -37,7 +38,7 @@ noncomputable section open TopologicalSpace MeasureTheory Filter -open scoped NNReal ENNReal Topology +open scoped NNReal ENNReal Topology ComplexConjugate variable {α E F G : Type*} {m m0 : MeasurableSpace α} {p : ℝ≥0∞} {q : ℝ} {μ ν : Measure α} [NormedAddCommGroup E] [NormedAddCommGroup F] [NormedAddCommGroup G] @@ -304,19 +305,23 @@ end Zero section Neg @[simp] -theorem eLpNorm'_neg {f : α → F} : eLpNorm' (-f) q μ = eLpNorm' f q μ := by simp [eLpNorm'] +theorem eLpNorm'_neg (f : α → F) (q : ℝ) (μ : Measure α) : eLpNorm' (-f) q μ = eLpNorm' f q μ := by + simp [eLpNorm'] @[deprecated (since := "2024-07-27")] alias snorm'_neg := eLpNorm'_neg @[simp] -theorem eLpNorm_neg {f : α → F} : eLpNorm (-f) p μ = eLpNorm f p μ := by +theorem eLpNorm_neg (f : α → F) (p : ℝ≥0∞) (μ : Measure α) : eLpNorm (-f) p μ = eLpNorm f p μ := by by_cases h0 : p = 0 · simp [h0] by_cases h_top : p = ∞ · simp [h_top, eLpNormEssSup] simp [eLpNorm_eq_eLpNorm' h0 h_top] +lemma eLpNorm_sub_comm (f g : α → E) (p : ℝ≥0∞) (μ : Measure α) : + eLpNorm (f - g) p μ = eLpNorm (g - f) p μ := by simp [← eLpNorm_neg (f := f - g)] + @[deprecated (since := "2024-07-27")] alias snorm_neg := eLpNorm_neg @@ -446,6 +451,8 @@ theorem memℒp_const_iff {p : ℝ≥0∞} {c : E} (hp_ne_zero : p ≠ 0) (hp_ne end Const +variable {f : α → F} + lemma eLpNorm'_mono_nnnorm_ae {f : α → F} {g : α → G} (hq : 0 ≤ q) (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ ‖g x‖₊) : eLpNorm' f q μ ≤ eLpNorm' g q μ := by simp only [eLpNorm'] @@ -828,8 +835,9 @@ private theorem eLpNorm_smul_measure_of_ne_zero_of_ne_top {p : ℝ≥0∞} (hp_n @[deprecated (since := "2024-07-27")] alias snorm_smul_measure_of_ne_zero_of_ne_top := eLpNorm_smul_measure_of_ne_zero_of_ne_top -theorem eLpNorm_smul_measure_of_ne_zero {p : ℝ≥0∞} {f : α → F} {c : ℝ≥0∞} (hc : c ≠ 0) : - eLpNorm f p (c • μ) = c ^ (1 / p).toReal • eLpNorm f p μ := by +/-- See `eLpNorm_smul_measure_of_ne_zero'` for a version with scalar multiplication by `ℝ≥0`. -/ +theorem eLpNorm_smul_measure_of_ne_zero {c : ℝ≥0∞} (hc : c ≠ 0) (f : α → F) (p : ℝ≥0∞) + (μ : Measure α) : eLpNorm f p (c • μ) = c ^ (1 / p).toReal • eLpNorm f p μ := by by_cases hp0 : p = 0 · simp [hp0] by_cases hp_top : p = ∞ @@ -839,12 +847,24 @@ theorem eLpNorm_smul_measure_of_ne_zero {p : ℝ≥0∞} {f : α → F} {c : ℝ @[deprecated (since := "2024-07-27")] alias snorm_smul_measure_of_ne_zero := eLpNorm_smul_measure_of_ne_zero -theorem eLpNorm_smul_measure_of_ne_top {p : ℝ≥0∞} (hp_ne_top : p ≠ ∞) {f : α → F} (c : ℝ≥0∞) : +/-- See `eLpNorm_smul_measure_of_ne_zero` for a version with scalar multiplication by `ℝ≥0∞`. -/ +lemma eLpNorm_smul_measure_of_ne_zero' {c : ℝ≥0} (hc : c ≠ 0) (f : α → F) (p : ℝ≥0∞) + (μ : Measure α) : eLpNorm f p (c • μ) = c ^ p.toReal⁻¹ • eLpNorm f p μ := + (eLpNorm_smul_measure_of_ne_zero (ENNReal.coe_ne_zero.2 hc) ..).trans (by simp; norm_cast) + +/-- See `eLpNorm_smul_measure_of_ne_top'` for a version with scalar multiplication by `ℝ≥0`. -/ +theorem eLpNorm_smul_measure_of_ne_top {p : ℝ≥0∞} (hp_ne_top : p ≠ ∞) (f : α → F) (c : ℝ≥0∞) : eLpNorm f p (c • μ) = c ^ (1 / p).toReal • eLpNorm f p μ := by by_cases hp0 : p = 0 · simp [hp0] · exact eLpNorm_smul_measure_of_ne_zero_of_ne_top hp0 hp_ne_top c +/-- See `eLpNorm_smul_measure_of_ne_top'` for a version with scalar multiplication by `ℝ≥0∞`. -/ +lemma eLpNorm_smul_measure_of_ne_top' (hp : p ≠ ∞) (c : ℝ≥0) (f : α → F) : + eLpNorm f p (c • μ) = c ^ p.toReal⁻¹ • eLpNorm f p μ := by + have : 0 ≤ p.toReal⁻¹ := by positivity + refine (eLpNorm_smul_measure_of_ne_top hp ..).trans ?_ + simp [ENNReal.smul_def, ENNReal.coe_rpow_of_nonneg, this] @[deprecated (since := "2024-07-27")] alias snorm_smul_measure_of_ne_top := eLpNorm_smul_measure_of_ne_top @@ -987,12 +1007,37 @@ theorem ae_le_eLpNormEssSup {f : α → F} : ∀ᵐ y ∂μ, ‖f y‖₊ ≤ eL @[deprecated (since := "2024-07-27")] alias ae_le_snormEssSup := ae_le_eLpNormEssSup +lemma eLpNormEssSup_lt_top_iff_isBoundedUnder : + eLpNormEssSup f μ < ⊤ ↔ IsBoundedUnder (· ≤ ·) (ae μ) fun x ↦ ‖f x‖₊ where + mp h := ⟨(eLpNormEssSup f μ).toNNReal, by + simp_rw [← ENNReal.coe_le_coe, ENNReal.coe_toNNReal h.ne]; exact ae_le_eLpNormEssSup⟩ + mpr := by rintro ⟨C, hC⟩; exact eLpNormEssSup_lt_top_of_ae_nnnorm_bound (C := C) hC + theorem meas_eLpNormEssSup_lt {f : α → F} : μ { y | eLpNormEssSup f μ < ‖f y‖₊ } = 0 := meas_essSup_lt @[deprecated (since := "2024-07-27")] alias meas_snormEssSup_lt := meas_eLpNormEssSup_lt +lemma eLpNorm_lt_top_of_finite [Finite α] [IsFiniteMeasure μ] : eLpNorm f p μ < ∞ := by + obtain rfl | hp₀ := eq_or_ne p 0 + · simp + obtain rfl | hp := eq_or_ne p ∞ + · simp only [eLpNorm_exponent_top, eLpNormEssSup_lt_top_iff_isBoundedUnder] + exact .le_of_finite + rw [eLpNorm_lt_top_iff_lintegral_rpow_nnnorm_lt_top hp₀ hp] + refine IsFiniteMeasure.lintegral_lt_top_of_bounded_to_ennreal μ ?_ + simp_rw [← ENNReal.coe_rpow_of_nonneg _ ENNReal.toReal_nonneg] + norm_cast + exact Finite.exists_le _ + +@[simp] lemma Memℒp.of_discrete [DiscreteMeasurableSpace α] [Finite α] [IsFiniteMeasure μ] : + Memℒp f p μ := + let ⟨C, hC⟩ := Finite.exists_le (‖f ·‖₊); .of_bound .of_finite C <| .of_forall hC + +@[simp] lemma eLpNorm_of_isEmpty [IsEmpty α] (f : α → E) (p : ℝ≥0∞) : eLpNorm f p μ = 0 := by + simp [Subsingleton.elim f 0] + lemma eLpNormEssSup_piecewise {s : Set α} (f g : α → E) [DecidablePred (· ∈ s)] (hs : MeasurableSet s) : eLpNormEssSup (Set.piecewise s f g) μ @@ -1196,37 +1241,33 @@ In this section we show inequalities on the norm. section BoundedSMul variable {𝕜 : Type*} [NormedRing 𝕜] [MulActionWithZero 𝕜 E] [MulActionWithZero 𝕜 F] -variable [BoundedSMul 𝕜 E] [BoundedSMul 𝕜 F] +variable [BoundedSMul 𝕜 E] [BoundedSMul 𝕜 F] {c : 𝕜} {f : α → F} -theorem eLpNorm'_const_smul_le (c : 𝕜) (f : α → F) (hq_pos : 0 < q) : - eLpNorm' (c • f) q μ ≤ ‖c‖₊ • eLpNorm' f q μ := - eLpNorm'_le_nnreal_smul_eLpNorm'_of_ae_le_mul (Eventually.of_forall fun _ => nnnorm_smul_le _ _) - hq_pos +theorem eLpNorm'_const_smul_le (hq : 0 < q) : eLpNorm' (c • f) q μ ≤ ‖c‖₊ • eLpNorm' f q μ := + eLpNorm'_le_nnreal_smul_eLpNorm'_of_ae_le_mul (Eventually.of_forall fun _ => nnnorm_smul_le ..) hq @[deprecated (since := "2024-07-27")] alias snorm'_const_smul_le := eLpNorm'_const_smul_le -theorem eLpNormEssSup_const_smul_le (c : 𝕜) (f : α → F) : - eLpNormEssSup (c • f) μ ≤ ‖c‖₊ • eLpNormEssSup f μ := +theorem eLpNormEssSup_const_smul_le : eLpNormEssSup (c • f) μ ≤ ‖c‖₊ • eLpNormEssSup f μ := eLpNormEssSup_le_nnreal_smul_eLpNormEssSup_of_ae_le_mul (Eventually.of_forall fun _ => by simp [nnnorm_smul_le]) @[deprecated (since := "2024-07-27")] alias snormEssSup_const_smul_le := eLpNormEssSup_const_smul_le -theorem eLpNorm_const_smul_le (c : 𝕜) (f : α → F) : eLpNorm (c • f) p μ ≤ ‖c‖₊ • eLpNorm f p μ := +theorem eLpNorm_const_smul_le : eLpNorm (c • f) p μ ≤ ‖c‖₊ • eLpNorm f p μ := eLpNorm_le_nnreal_smul_eLpNorm_of_ae_le_mul (Eventually.of_forall fun _ => by simp [nnnorm_smul_le]) _ @[deprecated (since := "2024-07-27")] alias snorm_const_smul_le := eLpNorm_const_smul_le -theorem Memℒp.const_smul {f : α → E} (hf : Memℒp f p μ) (c : 𝕜) : Memℒp (c • f) p μ := +theorem Memℒp.const_smul (hf : Memℒp f p μ) (c : 𝕜) : Memℒp (c • f) p μ := ⟨AEStronglyMeasurable.const_smul hf.1 c, - (eLpNorm_const_smul_le c f).trans_lt (ENNReal.mul_lt_top ENNReal.coe_lt_top hf.2)⟩ + eLpNorm_const_smul_le.trans_lt (ENNReal.mul_lt_top ENNReal.coe_lt_top hf.2)⟩ -theorem Memℒp.const_mul {R} [NormedRing R] {f : α → R} (hf : Memℒp f p μ) (c : R) : - Memℒp (fun x => c * f x) p μ := +theorem Memℒp.const_mul {f : α → 𝕜} (hf : Memℒp f p μ) (c : 𝕜) : Memℒp (fun x => c * f x) p μ := hf.const_smul c end BoundedSMul @@ -1246,9 +1287,8 @@ theorem eLpNorm'_const_smul {f : α → F} (c : 𝕜) (hq_pos : 0 < q) : eLpNorm' (c • f) q μ = ‖c‖₊ • eLpNorm' f q μ := by obtain rfl | hc := eq_or_ne c 0 · simp [eLpNorm', hq_pos] - refine le_antisymm (eLpNorm'_const_smul_le _ _ hq_pos) ?_ - have : eLpNorm' _ q μ ≤ _ := eLpNorm'_const_smul_le c⁻¹ (c • f) hq_pos - rwa [inv_smul_smul₀ hc, nnnorm_inv, le_inv_smul_iff_of_pos (nnnorm_pos.2 hc)] at this + refine le_antisymm (eLpNorm'_const_smul_le hq_pos) ?_ + simpa [hc, le_inv_smul_iff_of_pos] using eLpNorm'_const_smul_le (c := c⁻¹) (f := c • f) hq_pos @[deprecated (since := "2024-07-27")] alias snorm'_const_smul := eLpNorm'_const_smul @@ -1260,17 +1300,20 @@ theorem eLpNormEssSup_const_smul (c : 𝕜) (f : α → F) : @[deprecated (since := "2024-07-27")] alias snormEssSup_const_smul := eLpNormEssSup_const_smul -theorem eLpNorm_const_smul (c : 𝕜) (f : α → F) : +theorem eLpNorm_const_smul (c : 𝕜) (f : α → F) (p : ℝ≥0∞) (μ : Measure α): eLpNorm (c • f) p μ = (‖c‖₊ : ℝ≥0∞) * eLpNorm f p μ := by obtain rfl | hc := eq_or_ne c 0 · simp - refine le_antisymm (eLpNorm_const_smul_le _ _) ?_ - have : eLpNorm _ p μ ≤ _ := eLpNorm_const_smul_le c⁻¹ (c • f) - rwa [inv_smul_smul₀ hc, nnnorm_inv, le_inv_smul_iff_of_pos (nnnorm_pos.2 hc)] at this + refine le_antisymm eLpNorm_const_smul_le ?_ + simpa [hc, le_inv_smul_iff_of_pos] using eLpNorm_const_smul_le (c := c⁻¹) (f := c • f) @[deprecated (since := "2024-07-27")] alias snorm_const_smul := eLpNorm_const_smul +lemma eLpNorm_nsmul [NormedSpace ℝ F] (n : ℕ) (f : α → F) : + eLpNorm (n • f) p μ = n * eLpNorm f p μ := by + simpa [Nat.cast_smul_eq_nsmul] using eLpNorm_const_smul (n : ℝ) f .. + end NormedSpace theorem le_eLpNorm_of_bddBelow (hp : p ≠ 0) (hp' : p ≠ ∞) {f : α → F} (C : ℝ≥0) {s : Set α} @@ -1299,6 +1342,9 @@ section RCLike variable {𝕜 : Type*} [RCLike 𝕜] {f : α → 𝕜} +@[simp] lemma eLpNorm_conj (f : α → 𝕜) (p : ℝ≥0∞) (μ : Measure α) : + eLpNorm (conj f) p μ = eLpNorm f p μ := by simp [← eLpNorm_norm] + theorem Memℒp.re (hf : Memℒp f p μ) : Memℒp (fun x => RCLike.re (f x)) p μ := by have : ∀ x, ‖RCLike.re (f x)‖ ≤ 1 * ‖f x‖ := by intro x diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index 0d14ecf974a22..ca21a90b697cf 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -426,7 +426,7 @@ variable [BoundedSMul 𝕜 E] [BoundedSMul 𝕜' E] theorem const_smul_mem_Lp (c : 𝕜) (f : Lp E p μ) : c • (f : α →ₘ[μ] E) ∈ Lp E p μ := by rw [mem_Lp_iff_eLpNorm_lt_top, eLpNorm_congr_ae (AEEqFun.coeFn_smul _ _)] - refine (eLpNorm_const_smul_le _ _).trans_lt ?_ + refine eLpNorm_const_smul_le.trans_lt ?_ rw [ENNReal.smul_def, smul_eq_mul, ENNReal.mul_lt_top_iff] exact Or.inl ⟨ENNReal.coe_lt_top, f.prop⟩ @@ -464,7 +464,7 @@ instance instBoundedSMul [Fact (1 ≤ p)] : BoundedSMul 𝕜 (Lp E p μ) := suffices (‖r • f‖₊ : ℝ≥0∞) ≤ ‖r‖₊ * ‖f‖₊ from mod_cast this rw [nnnorm_def, nnnorm_def, ENNReal.coe_toNNReal (Lp.eLpNorm_ne_top _), eLpNorm_congr_ae (coeFn_smul _ _), ENNReal.coe_toNNReal (Lp.eLpNorm_ne_top _)] - exact eLpNorm_const_smul_le r f + exact eLpNorm_const_smul_le end BoundedSMul From 5ad34c033417c6e6efd3fcd9062fa1764d240209 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Tue, 26 Nov 2024 13:02:57 +0000 Subject: [PATCH 320/829] feat(Algebra/Exact): iff lemmas (#19504) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit When we have a commutative diagram from a sequence of two maps to another, such that the left vertical map is surjective, the middle vertical map is bijective and the right vertical map is injective, then the upper row is exact iff the lower row is. This generalizes the case when all the three vertical maps are bijections. This PR introduces two new lemmas `exact_iff_of_surjective_of_bijective_of_injective` in the `AddMonoidHom` and `LinearMap` namespaces, and refactors the proofs of the all-bijections versions. This parallels the categorical lemma https://leanprover-community.github.io/mathlib4_docs/Mathlib/Algebra/Homology/ShortComplex/Exact.html#CategoryTheory.ShortComplex.exact_iff_of_epi_of_isIso_of_mono Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- Mathlib/Algebra/Exact.lean | 96 ++++++++++++++++++++++++++------------ 1 file changed, 67 insertions(+), 29 deletions(-) diff --git a/Mathlib/Algebra/Exact.lean b/Mathlib/Algebra/Exact.lean index 13f90a7ac5eca..c9da18e29ffbc 100644 --- a/Mathlib/Algebra/Exact.lean +++ b/Mathlib/Algebra/Exact.lean @@ -90,6 +90,44 @@ lemma exact_of_comp_of_mem_range (h1 : g.comp f = 0) (h2 : ∀ x, g x = 0 → x ∈ range f) : Exact f g := exact_of_comp_eq_zero_of_ker_le_range h1 h2 +/-- When we have a commutative diagram from a sequence of two maps to another, +such that the left vertical map is surjective, the middle vertical map is bijective and the right +vertical map is injective, then the upper row is exact iff the lower row is. +See `ShortComplex.exact_iff_of_epi_of_isIso_of_mono` in the file +`Algebra.Homology.ShortComplex.Exact` for the categorical version of this result. -/ +lemma exact_iff_of_surjective_of_bijective_of_injective + {M₁ M₂ M₃ N₁ N₂ N₃ : Type*} [AddCommMonoid M₁] [AddCommMonoid M₂] [AddCommMonoid M₃] + [AddCommMonoid N₁] [AddCommMonoid N₂] [AddCommMonoid N₃] + (f : M₁ →+ M₂) (g : M₂ →+ M₃) (f' : N₁ →+ N₂) (g' : N₂ →+ N₃) + (τ₁ : M₁ →+ N₁) (τ₂ : M₂ →+ N₂) (τ₃ : M₃ →+ N₃) + (comm₁₂ : f'.comp τ₁ = τ₂.comp f) + (comm₂₃ : g'.comp τ₂ = τ₃.comp g) + (h₁ : Function.Surjective τ₁) (h₂ : Function.Bijective τ₂) (h₃ : Function.Injective τ₃) : + Exact f g ↔ Exact f' g' := by + replace comm₁₂ := DFunLike.congr_fun comm₁₂ + replace comm₂₃ := DFunLike.congr_fun comm₂₃ + dsimp at comm₁₂ comm₂₃ + constructor + · intro h y₂ + obtain ⟨x₂, rfl⟩ := h₂.2 y₂ + constructor + · intro hx₂ + obtain ⟨x₁, rfl⟩ := (h x₂).1 (h₃ (by simpa only [map_zero, comm₂₃] using hx₂)) + exact ⟨τ₁ x₁, by simp only [comm₁₂]⟩ + · rintro ⟨y₁, hy₁⟩ + obtain ⟨x₁, rfl⟩ := h₁ y₁ + rw [comm₂₃, (h x₂).2 _, map_zero] + exact ⟨x₁, h₂.1 (by simpa only [comm₁₂] using hy₁)⟩ + · intro h x₂ + constructor + · intro hx₂ + obtain ⟨y₁, hy₁⟩ := (h (τ₂ x₂)).1 (by simp only [comm₂₃, hx₂, map_zero]) + obtain ⟨x₁, rfl⟩ := h₁ y₁ + exact ⟨x₁, h₂.1 (by simpa only [comm₁₂] using hy₁)⟩ + · rintro ⟨x₁, rfl⟩ + apply h₃ + simp only [← comm₁₂, ← comm₂₃, h.apply_apply_eq_zero (τ₁ x₁), map_zero] + end AddMonoidHom namespace Function.Exact @@ -110,38 +148,18 @@ variable {X₁ X₂ X₃ Y₁ Y₂ Y₃ : Type*} [AddCommMonoid X₁] [AddCommMo (e₁ : X₁ ≃+ Y₁) (e₂ : X₂ ≃+ Y₂) (e₃ : X₃ ≃+ Y₃) {f₁₂ : X₁ →+ X₂} {f₂₃ : X₂ →+ X₃} {g₁₂ : Y₁ →+ Y₂} {g₂₃ : Y₂ →+ Y₃} +lemma iff_of_ladder_addEquiv (comm₁₂ : g₁₂.comp e₁ = AddMonoidHom.comp e₂ f₁₂) + (comm₂₃ : g₂₃.comp e₂ = AddMonoidHom.comp e₃ f₂₃) : Exact g₁₂ g₂₃ ↔ Exact f₁₂ f₂₃ := + (exact_iff_of_surjective_of_bijective_of_injective _ _ _ _ e₁ e₂ e₃ comm₁₂ comm₂₃ + e₁.surjective e₂.bijective e₃.injective).symm + lemma of_ladder_addEquiv_of_exact (comm₁₂ : g₁₂.comp e₁ = AddMonoidHom.comp e₂ f₁₂) - (comm₂₃ : g₂₃.comp e₂ = AddMonoidHom.comp e₃ f₂₃) (H : Exact f₁₂ f₂₃) : Exact g₁₂ g₂₃ := by - have h₁₂ := DFunLike.congr_fun comm₁₂ - have h₂₃ := DFunLike.congr_fun comm₂₃ - dsimp at h₁₂ h₂₃ - apply of_comp_eq_zero_of_ker_in_range - · ext y₁ - obtain ⟨x₁, rfl⟩ := e₁.surjective y₁ - dsimp - rw [h₁₂, h₂₃, H.apply_apply_eq_zero, map_zero] - · intro y₂ hx₂ - obtain ⟨x₂, rfl⟩ := e₂.surjective y₂ - obtain ⟨x₁, rfl⟩ := (H x₂).1 (e₃.injective (by rw [← h₂₃, hx₂, map_zero])) - exact ⟨e₁ x₁, by rw [h₁₂]⟩ + (comm₂₃ : g₂₃.comp e₂ = AddMonoidHom.comp e₃ f₂₃) (H : Exact f₁₂ f₂₃) : Exact g₁₂ g₂₃ := + (iff_of_ladder_addEquiv _ _ _ comm₁₂ comm₂₃).2 H lemma of_ladder_addEquiv_of_exact' (comm₁₂ : g₁₂.comp e₁ = AddMonoidHom.comp e₂ f₁₂) - (comm₂₃ : g₂₃.comp e₂ = AddMonoidHom.comp e₃ f₂₃) (H : Exact g₁₂ g₂₃) : Exact f₁₂ f₂₃ := by - refine of_ladder_addEquiv_of_exact e₁.symm e₂.symm e₃.symm ?_ ?_ H - · ext y₁ - obtain ⟨x₁, rfl⟩ := e₁.surjective y₁ - apply e₂.injective - simpa using DFunLike.congr_fun comm₁₂.symm x₁ - · ext y₂ - obtain ⟨x₂, rfl⟩ := e₂.surjective y₂ - apply e₃.injective - simpa using DFunLike.congr_fun comm₂₃.symm x₂ - -lemma iff_of_ladder_addEquiv (comm₁₂ : g₁₂.comp e₁ = AddMonoidHom.comp e₂ f₁₂) - (comm₂₃ : g₂₃.comp e₂ = AddMonoidHom.comp e₃ f₂₃) : Exact g₁₂ g₂₃ ↔ Exact f₁₂ f₂₃ := by - constructor - · exact of_ladder_addEquiv_of_exact' e₁ e₂ e₃ comm₁₂ comm₂₃ - · exact of_ladder_addEquiv_of_exact e₁ e₂ e₃ comm₁₂ comm₂₃ + (comm₂₃ : g₂₃.comp e₂ = AddMonoidHom.comp e₃ f₂₃) (H : Exact g₁₂ g₂₃) : Exact f₁₂ f₂₃ := + (iff_of_ladder_addEquiv _ _ _ comm₁₂ comm₂₃).1 H end @@ -396,6 +414,26 @@ end Function namespace LinearMap +/-- When we have a commutative diagram from a sequence of two linear maps to another, +such that the left vertical map is surjective, the middle vertical map is bijective and the right +vertical map is injective, then the upper row is exact iff the lower row is. +See `ShortComplex.exact_iff_of_epi_of_isIso_of_mono` in the file +`Algebra.Homology.ShortComplex.Exact` for the categorical version of this result. -/ +lemma exact_iff_of_surjective_of_bijective_of_injective + {M₁ M₂ M₃ N₁ N₂ N₃ : Type*} [AddCommMonoid M₁] [AddCommMonoid M₂] [AddCommMonoid M₃] + [AddCommMonoid N₁] [AddCommMonoid N₂] [AddCommMonoid N₃] + [Module R M₁] [Module R M₂] [Module R M₃] + [Module R N₁] [Module R N₂] [Module R N₃] + (f : M₁ →ₗ[R] M₂) (g : M₂ →ₗ[R] M₃) (f' : N₁ →ₗ[R] N₂) (g' : N₂ →ₗ[R] N₃) + (τ₁ : M₁ →ₗ[R] N₁) (τ₂ : M₂ →ₗ[R] N₂) (τ₃ : M₃ →ₗ[R] N₃) + (comm₁₂ : f'.comp τ₁ = τ₂.comp f) (comm₂₃ : g'.comp τ₂ = τ₃.comp g) + (h₁ : Function.Surjective τ₁) (h₂ : Function.Bijective τ₂) (h₃ : Function.Injective τ₃) : + Function.Exact f g ↔ Function.Exact f' g' := + AddMonoidHom.exact_iff_of_surjective_of_bijective_of_injective + f.toAddMonoidHom g.toAddMonoidHom f'.toAddMonoidHom g'.toAddMonoidHom + τ₁.toAddMonoidHom τ₂.toAddMonoidHom τ₃.toAddMonoidHom + (by ext; apply DFunLike.congr_fun comm₁₂) (by ext; apply DFunLike.congr_fun comm₂₃) h₁ h₂ h₃ + lemma surjective_range_liftQ (h : range f ≤ ker g) (hg : Function.Surjective g) : Function.Surjective ((range f).liftQ g h) := by intro x₃ From 13fc0e8a77fcb3fc528f47e1e5a1b09df301bb0c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 26 Nov 2024 14:07:40 +0000 Subject: [PATCH 321/829] feat(Pointwise): monotonicity of `pow` (#19252) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit ... and `s ^ n = ∅ ↔ s = ∅ ∧ n ≠ 0` From GrowthInGroups --- Mathlib/Algebra/Algebra/Operations.lean | 2 +- .../Algebra/Group/Pointwise/Finset/Basic.lean | 134 ++++++++++++++---- .../Algebra/Group/Pointwise/Set/Basic.lean | 121 ++++++++++++---- .../Algebra/Order/Monoid/Unbundled/Pow.lean | 14 ++ Mathlib/Analysis/Convex/EGauge.lean | 3 +- Mathlib/GroupTheory/OrderOfElement.lean | 2 +- 6 files changed, 214 insertions(+), 62 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Operations.lean b/Mathlib/Algebra/Algebra/Operations.lean index 960f805fd05b7..7c503761b174e 100644 --- a/Mathlib/Algebra/Algebra/Operations.lean +++ b/Mathlib/Algebra/Algebra/Operations.lean @@ -230,7 +230,7 @@ theorem pow_subset_pow {n : ℕ} : (↑M : Set A) ^ n ⊆ ↑(M ^ n : Submodule trans AddSubmonoid.pow_subset_pow (le_pow_toAddSubmonoid M) theorem pow_mem_pow {x : A} (hx : x ∈ M) (n : ℕ) : x ^ n ∈ M ^ n := - pow_subset_pow _ <| Set.pow_mem_pow hx _ + pow_subset_pow _ <| Set.pow_mem_pow hx end Module diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index ac78f4f5b38b5..876f75a97121b 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -544,14 +544,17 @@ theorem singleton_mul_singleton (a b : α) : ({a} : Finset α) * {b} = {a * b} : theorem mul_subset_mul : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ * t₁ ⊆ s₂ * t₂ := image₂_subset -@[to_additive] +@[to_additive (attr := gcongr)] theorem mul_subset_mul_left : t₁ ⊆ t₂ → s * t₁ ⊆ s * t₂ := image₂_subset_left -@[to_additive] +@[to_additive (attr := gcongr)] theorem mul_subset_mul_right : s₁ ⊆ s₂ → s₁ * t ⊆ s₂ * t := image₂_subset_right +@[to_additive] instance : MulLeftMono (Finset α) where elim _s _t₁ _t₂ := mul_subset_mul_left +@[to_additive] instance : MulRightMono (Finset α) where elim _t _s₁ _s₂ := mul_subset_mul_right + @[to_additive] theorem mul_subset_iff : s * t ⊆ u ↔ ∀ x ∈ s, ∀ y ∈ t, x * y ∈ u := image₂_subset_iff @@ -947,29 +950,69 @@ protected def monoid : Monoid (Finset α) := scoped[Pointwise] attribute [instance] Finset.monoid Finset.addMonoid +-- `Finset.pow_left_monotone` doesn't exist since it would syntactically be a special case of +-- `pow_left_mono` + @[to_additive] -theorem pow_mem_pow (ha : a ∈ s) : ∀ n : ℕ, a ^ n ∈ s ^ n - | 0 => by - simp only [pow_zero, mem_one] - | n + 1 => by - simp only [pow_succ] - exact mul_mem_mul (pow_mem_pow ha n) ha +protected lemma pow_right_monotone (hs : 1 ∈ s) : Monotone (s ^ ·) := + pow_right_monotone <| one_subset.2 hs + +@[to_additive (attr := gcongr)] +lemma pow_subset_pow_left (hst : s ⊆ t) : s ^ n ⊆ t ^ n := pow_left_mono _ hst + +@[to_additive (attr := gcongr)] +lemma pow_subset_pow_right (hs : 1 ∈ s) (hmn : m ≤ n) : s ^ m ⊆ s ^ n := + Finset.pow_right_monotone hs hmn + +@[to_additive (attr := gcongr)] +lemma pow_subset_pow (hst : s ⊆ t) (ht : 1 ∈ t) (hmn : m ≤ n) : s ^ m ⊆ t ^ n := + (pow_subset_pow_left hst).trans (pow_subset_pow_right ht hmn) + +@[deprecated (since := "2024-11-19")] alias pow_subset_pow_of_one_mem := pow_subset_pow_right + +@[deprecated (since := "2024-11-19")] +alias nsmul_subset_nsmul_of_zero_mem := nsmul_subset_nsmul_right + +@[to_additive] +lemma pow_subset_pow_mul_of_sq_subset_mul (hst : s ^ 2 ⊆ t * s) (hn : n ≠ 0) : + s ^ n ⊆ t ^ (n - 1) * s := pow_le_pow_mul_of_sq_le_mul hst hn + +@[to_additive (attr := simp) nsmul_empty] +lemma empty_pow (hn : n ≠ 0) : (∅ : Finset α) ^ n = ∅ := match n with | n + 1 => by simp [pow_succ] + +@[deprecated (since := "2024-10-21")] alias empty_nsmul := nsmul_empty @[to_additive] -theorem pow_subset_pow (hst : s ⊆ t) : ∀ n : ℕ, s ^ n ⊆ t ^ n - | 0 => by - simp [pow_zero] - | n + 1 => by - rw [pow_succ] - exact mul_subset_mul (pow_subset_pow hst n) hst +lemma Nonempty.pow (hs : s.Nonempty) : ∀ {n}, (s ^ n).Nonempty + | 0 => by simp + | n + 1 => by rw [pow_succ]; exact hs.pow.mul hs + +set_option push_neg.use_distrib true in +@[to_additive (attr := simp)] lemma pow_eq_empty : s ^ n = ∅ ↔ s = ∅ ∧ n ≠ 0 := by + constructor + · contrapose! + rintro (hs | rfl) + -- TODO: The `nonempty_iff_ne_empty` would be unnecessary if `push_neg` knew how to simplify + -- `s ≠ ∅` to `s.Nonempty` when `s : Finset α`. + -- See https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/push_neg.20extensibility + · exact nonempty_iff_ne_empty.1 (nonempty_iff_ne_empty.2 hs).pow + · rw [← nonempty_iff_ne_empty] + simp + · rintro ⟨rfl, hn⟩ + exact empty_pow hn + +@[to_additive (attr := simp) nsmul_singleton] +lemma singleton_pow (a : α) : ∀ n, ({a} : Finset α) ^ n = {a ^ n} + | 0 => by simp [singleton_one] + | n + 1 => by simp [pow_succ, singleton_pow _ n] + +@[to_additive] lemma pow_mem_pow (ha : a ∈ s) : a ^ n ∈ s ^ n := by + simpa using pow_subset_pow_left (singleton_subset_iff.2 ha) + +@[to_additive] lemma one_mem_pow (hs : 1 ∈ s) : 1 ∈ s ^ n := by simpa using pow_mem_pow hs @[to_additive] -theorem pow_subset_pow_of_one_mem (hs : (1 : α) ∈ s) : m ≤ n → s ^ m ⊆ s ^ n := by - apply Nat.le_induction - · exact fun _ hn => hn - · intro n _ hmn - rw [pow_succ] - exact hmn.trans (subset_mul_left (s ^ n) hs) +lemma inter_pow_subset : (s ∩ t) ^ n ⊆ s ^ n ∩ t ^ n := by apply subset_inter <;> gcongr <;> simp @[to_additive (attr := simp, norm_cast)] theorem coe_list_prod (s : List (Finset α)) : (↑s.prod : Set α) = (s.map (↑)).prod := @@ -986,17 +1029,6 @@ theorem mem_pow {a : α} {n : ℕ} : a ∈ s ^ n ↔ ∃ f : Fin n → s, (List.ofFn fun i => ↑(f i)).prod = a := by simp [← mem_coe, coe_pow, Set.mem_pow] -@[to_additive (attr := simp) nsmul_empty] -theorem empty_pow (hn : n ≠ 0) : (∅ : Finset α) ^ n = ∅ := by - rw [← tsub_add_cancel_of_le (Nat.succ_le_of_lt <| Nat.pos_of_ne_zero hn), pow_succ', empty_mul] - -@[deprecated (since := "2024-10-21")] alias empty_nsmul := nsmul_empty - -@[to_additive (attr := simp) nsmul_singleton] -lemma singleton_pow (a : α) : ∀ n, ({a} : Finset α) ^ n = {a ^ n} - | 0 => by simp [singleton_one] - | n + 1 => by simp [pow_succ, singleton_pow _ n] - @[to_additive] lemma card_pow_le : ∀ {n}, (s ^ n).card ≤ s.card ^ n | 0 => by simp @@ -1105,6 +1137,22 @@ lemma univ_div_univ [Fintype α] : (univ / univ : Finset α) = univ := by simp [ @[to_additive (attr := simp) zsmul_empty] lemma empty_zpow (hn : n ≠ 0) : (∅ : Finset α) ^ n = ∅ := by cases n <;> aesop +@[to_additive] +lemma Nonempty.zpow (hs : s.Nonempty) : ∀ {n : ℤ}, (s ^ n).Nonempty + | (n : ℕ) => hs.pow + | .negSucc n => by simpa using hs.pow + +set_option push_neg.use_distrib true in +@[to_additive (attr := simp)] lemma zpow_eq_empty : s ^ n = ∅ ↔ s = ∅ ∧ n ≠ 0 := by + constructor + · contrapose! + rintro (hs | rfl) + · exact nonempty_iff_ne_empty.1 (nonempty_iff_ne_empty.2 hs).zpow + · rw [← nonempty_iff_ne_empty] + simp + · rintro ⟨rfl, hn⟩ + exact empty_zpow hn + @[to_additive (attr := simp) zsmul_singleton] lemma singleton_zpow (a : α) (n : ℤ) : ({a} : Finset α) ^ n = {a ^ n} := by cases n <;> simp @@ -1129,10 +1177,18 @@ variable (f : F) {s t : Finset α} {a b : α} theorem one_mem_div_iff : (1 : α) ∈ s / t ↔ ¬Disjoint s t := by rw [← mem_coe, ← disjoint_coe, coe_div, Set.one_mem_div_iff] +@[to_additive (attr := simp)] +lemma one_mem_inv_mul_iff : (1 : α) ∈ t⁻¹ * s ↔ ¬Disjoint s t := by + aesop (add simp [not_disjoint_iff_nonempty_inter, mem_mul, mul_eq_one_iff_eq_inv, + Finset.Nonempty]) + @[to_additive] theorem not_one_mem_div_iff : (1 : α) ∉ s / t ↔ Disjoint s t := one_mem_div_iff.not_left +@[to_additive] +lemma not_one_mem_inv_mul_iff : (1 : α) ∉ t⁻¹ * s ↔ Disjoint s t := one_mem_inv_mul_iff.not_left + @[to_additive] theorem Nonempty.one_mem_div (h : s.Nonempty) : (1 : α) ∈ s / s := let ⟨a, ha⟩ := h @@ -1430,6 +1486,22 @@ theorem mul_subset_iff_right : s * t ⊆ u ↔ ∀ b ∈ t, op b • s ⊆ u := end Mul +section Monoid +variable [DecidableEq α] [DecidableEq β] [Monoid α] [Monoid β] [FunLike F α β] + +@[to_additive] +lemma image_pow_of_ne_zero [MulHomClass F α β] : + ∀ {n}, n ≠ 0 → ∀ (f : F) (s : Finset α), (s ^ n).image f = s.image f ^ n + | 1, _ => by simp + | n + 2, _ => by simp [image_mul, pow_succ _ n.succ, image_pow_of_ne_zero] + +@[to_additive] +lemma image_pow [MonoidHomClass F α β] (f : F) (s : Finset α) : ∀ n, (s ^ n).image f = s.image f ^ n + | 0 => by simp [singleton_one] + | n + 1 => image_pow_of_ne_zero n.succ_ne_zero .. + +end Monoid + section Semigroup variable [Semigroup α] [DecidableEq α] diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean index 417727b33b985..a834f509b4e16 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean @@ -6,6 +6,7 @@ Authors: Johan Commelin, Floris van Doorn import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.Algebra.Group.Units.Hom import Mathlib.Algebra.Opposites +import Mathlib.Algebra.Order.Monoid.Unbundled.Pow import Mathlib.Data.Set.Lattice /-! @@ -106,7 +107,7 @@ theorem one_mem_one : (1 : α) ∈ (1 : Set α) := theorem one_subset : 1 ⊆ s ↔ (1 : α) ∈ s := singleton_subset_iff -@[to_additive] +@[to_additive (attr := simp)] theorem one_nonempty : (1 : Set α).Nonempty := ⟨1, rfl⟩ @@ -347,6 +348,9 @@ theorem mul_subset_mul_left : t₁ ⊆ t₂ → s * t₁ ⊆ s * t₂ := theorem mul_subset_mul_right : s₁ ⊆ s₂ → s₁ * t ⊆ s₂ * t := image2_subset_right +@[to_additive] instance : MulLeftMono (Set α) where elim _s _t₁ _t₂ := mul_subset_mul_left +@[to_additive] instance : MulRightMono (Set α) where elim _t _s₁ _s₂ := mul_subset_mul_right + @[to_additive] theorem mul_subset_iff : s * t ⊆ u ↔ ∀ x ∈ s, ∀ y ∈ t, x * y ∈ u := image2_subset_iff @@ -1075,46 +1079,66 @@ protected noncomputable def monoid : Monoid (Set α) := scoped[Pointwise] attribute [instance] Set.monoid Set.addMonoid -@[to_additive] -theorem pow_mem_pow (ha : a ∈ s) : ∀ n : ℕ, a ^ n ∈ s ^ n - | 0 => by - simp only [pow_zero, mem_one] - | n + 1 => by - simp only [pow_succ] - exact mul_mem_mul (pow_mem_pow ha _) ha +-- `Set.pow_left_monotone` doesn't exist since it would syntactically be a special case of +-- `pow_left_mono` @[to_additive] -theorem pow_subset_pow (hst : s ⊆ t) : ∀ n : ℕ, s ^ n ⊆ t ^ n - | 0 => by - rw [pow_zero] - exact Subset.rfl - | n + 1 => by - rw [pow_succ] - exact mul_subset_mul (pow_subset_pow hst _) hst +protected lemma pow_right_monotone (hs : 1 ∈ s) : Monotone (s ^ ·) := + pow_right_monotone <| one_subset.2 hs + +@[to_additive (attr := gcongr)] +lemma pow_subset_pow_left (hst : s ⊆ t) : s ^ n ⊆ t ^ n := pow_left_mono _ hst + +@[to_additive (attr := gcongr)] +lemma pow_subset_pow_right (hs : 1 ∈ s) (hmn : m ≤ n) : s ^ m ⊆ s ^ n := + Set.pow_right_monotone hs hmn + +@[to_additive (attr := gcongr)] +lemma pow_subset_pow (hst : s ⊆ t) (ht : 1 ∈ t) (hmn : m ≤ n) : s ^ m ⊆ t ^ n := + (pow_subset_pow_left hst).trans (pow_subset_pow_right ht hmn) + +@[deprecated (since := "2024-11-19")] alias pow_subset_pow_of_one_mem := pow_subset_pow_right + +@[deprecated (since := "2024-11-19")] +alias nsmul_subset_nsmul_of_zero_mem := nsmul_subset_nsmul_right @[to_additive] -theorem pow_subset_pow_of_one_mem (hs : (1 : α) ∈ s) (hn : m ≤ n) : s ^ m ⊆ s ^ n := by - -- Porting note: `Nat.le_induction` didn't work as an induction principle in mathlib3, this was - -- `refine Nat.le_induction ...` - induction n, hn using Nat.le_induction with - | base => exact Subset.rfl - | succ _ _ ih => - dsimp only - rw [pow_succ'] - exact ih.trans (subset_mul_right _ hs) +lemma pow_subset_pow_mul_of_sq_subset_mul (hst : s ^ 2 ⊆ t * s) (hn : n ≠ 0) : + s ^ n ⊆ t ^ (n - 1) * s := pow_le_pow_mul_of_sq_le_mul hst hn @[to_additive (attr := simp) nsmul_empty] -theorem empty_pow {n : ℕ} (hn : n ≠ 0) : (∅ : Set α) ^ n = ∅ := by - match n with - | n + 1 => rw [pow_succ', empty_mul] +lemma empty_pow (hn : n ≠ 0) : (∅ : Set α) ^ n = ∅ := match n with | n + 1 => by simp [pow_succ] @[deprecated (since := "2024-10-21")] alias empty_nsmul := nsmul_empty +@[to_additive] +lemma Nonempty.pow (hs : s.Nonempty) : ∀ {n}, (s ^ n).Nonempty + | 0 => by simp + | n + 1 => by rw [pow_succ]; exact hs.pow.mul hs + +set_option push_neg.use_distrib true in +@[to_additive (attr := simp)] lemma pow_eq_empty : s ^ n = ∅ ↔ s = ∅ ∧ n ≠ 0 := by + constructor + · contrapose! + rintro (hs | rfl) + · exact hs.pow + · simp + · rintro ⟨rfl, hn⟩ + exact empty_pow hn + @[to_additive (attr := simp) nsmul_singleton] lemma singleton_pow (a : α) : ∀ n, ({a} : Set α) ^ n = {a ^ n} | 0 => by simp [singleton_one] | n + 1 => by simp [pow_succ, singleton_pow _ n] +@[to_additive] lemma pow_mem_pow (ha : a ∈ s) : a ^ n ∈ s ^ n := by + simpa using pow_subset_pow_left (singleton_subset_iff.2 ha) + +@[to_additive] lemma one_mem_pow (hs : 1 ∈ s) : 1 ∈ s ^ n := by simpa using pow_mem_pow hs + +@[to_additive] +lemma inter_pow_subset : (s ∩ t) ^ n ⊆ s ^ n ∩ t ^ n := by apply subset_inter <;> gcongr <;> simp + @[to_additive] theorem mul_univ_of_one_mem (hs : (1 : α) ∈ s) : s * univ = univ := eq_univ_iff_forall.2 fun _ => mem_mul.2 ⟨_, hs, _, mem_univ _, one_mul _⟩ @@ -1206,6 +1230,21 @@ lemma univ_div_univ : (univ / univ : Set α) = univ := by simp [div_eq_mul_inv] @[to_additive (attr := simp) zsmul_empty] lemma empty_zpow (hn : n ≠ 0) : (∅ : Set α) ^ n = ∅ := by cases n <;> aesop +@[to_additive] +lemma Nonempty.zpow (hs : s.Nonempty) : ∀ {n : ℤ}, (s ^ n).Nonempty + | (n : ℕ) => hs.pow + | .negSucc n => by simpa using hs.pow + +set_option push_neg.use_distrib true in +@[to_additive (attr := simp)] lemma zpow_eq_empty : s ^ n = ∅ ↔ s = ∅ ∧ n ≠ 0 := by + constructor + · contrapose! + rintro (hs | rfl) + · exact hs.zpow + · simp + · rintro ⟨rfl, hn⟩ + exact empty_zpow hn + @[to_additive (attr := simp) zsmul_singleton] lemma singleton_zpow (a : α) (n : ℤ) : ({a} : Set α) ^ n = {a ^ n} := by cases n <;> simp @@ -1231,10 +1270,17 @@ variable [Group α] {s t : Set α} {a b : α} theorem one_mem_div_iff : (1 : α) ∈ s / t ↔ ¬Disjoint s t := by simp [not_disjoint_iff_nonempty_inter, mem_div, div_eq_one, Set.Nonempty] +@[to_additive (attr := simp)] +lemma one_mem_inv_mul_iff : (1 : α) ∈ t⁻¹ * s ↔ ¬Disjoint s t := by + aesop (add simp [not_disjoint_iff_nonempty_inter, mem_mul, mul_eq_one_iff_eq_inv, Set.Nonempty]) + @[to_additive] theorem not_one_mem_div_iff : (1 : α) ∉ s / t ↔ Disjoint s t := one_mem_div_iff.not_left +@[to_additive] +lemma not_one_mem_inv_mul_iff : (1 : α) ∉ t⁻¹ * s ↔ Disjoint s t := one_mem_inv_mul_iff.not_left + alias ⟨_, _root_.Disjoint.one_not_mem_div_set⟩ := not_one_mem_div_iff attribute [to_additive] Disjoint.one_not_mem_div_set @@ -1334,6 +1380,22 @@ lemma preimage_mul (hm : Injective m) {s t : Set β} (hs : s ⊆ range m) (ht : end Mul +section Monoid +variable [Monoid α] [Monoid β] [FunLike F α β] + +@[to_additive] +lemma image_pow_of_ne_zero [MulHomClass F α β] : + ∀ {n}, n ≠ 0 → ∀ (f : F) (s : Set α), f '' (s ^ n) = (f '' s) ^ n + | 1, _ => by simp + | n + 2, _ => by simp [image_mul, pow_succ _ n.succ, image_pow_of_ne_zero] + +@[to_additive] +lemma image_pow [MonoidHomClass F α β] (f : F) (s : Set α) : ∀ n, f '' (s ^ n) = (f '' s) ^ n + | 0 => by simp [singleton_one] + | n + 1 => image_pow_of_ne_zero n.succ_ne_zero .. + +end Monoid + section Group variable [Group α] [DivisionMonoid β] [FunLike F α β] [MonoidHomClass F α β] (m : F) {s t : Set α} @@ -1363,6 +1425,11 @@ lemma preimage_div (hm : Injective m) {s t : Set β} (hs : s ⊆ range m) (ht : end Group +variable {ι : Type*} {α : ι → Type*} [∀ i, Inv (α i)] + +@[to_additive (attr := simp)] +lemma inv_pi (s : Set ι) (t : ∀ i, Set (α i)) : (s.pi t)⁻¹ = s.pi fun i ↦ (t i)⁻¹ := by ext x; simp + end Set /-! ### Miscellaneous -/ diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean index 0a60e09d11d81..32f0ddb84013a 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean @@ -119,6 +119,20 @@ theorem Right.pow_lt_one_of_lt {n : ℕ} {x : M} (hn : 0 < n) (h : x < 1) : x ^ @[deprecated (since := "2024-09-21")] alias Right.pow_neg := Right.nsmul_neg +/-- This lemma is useful in non-cancellative monoids, like sets under pointwise operations. -/ +@[to_additive +"This lemma is useful in non-cancellative monoids, like sets under pointwise operations."] +lemma pow_le_pow_mul_of_sq_le_mul [MulLeftMono M] {a b : M} (hab : a ^ 2 ≤ b * a) : + ∀ {n}, n ≠ 0 → a ^ n ≤ b ^ (n - 1) * a + | 1, _ => by simp + | n + 2, _ => by + calc + a ^ (n + 2) = a ^ (n + 1) * a := by rw [pow_succ] + _ ≤ b ^ n * a * a := mul_le_mul_right' (pow_le_pow_mul_of_sq_le_mul hab (by omega)) _ + _ = b ^ n * a ^ 2 := by rw [mul_assoc, sq] + _ ≤ b ^ n * (b * a) := mul_le_mul_left' hab _ + _ = b ^ (n + 1) * a := by rw [← mul_assoc, ← pow_succ] + end Right section CovariantLTSwap diff --git a/Mathlib/Analysis/Convex/EGauge.lean b/Mathlib/Analysis/Convex/EGauge.lean index 7ade295422705..6a61629fae9e7 100644 --- a/Mathlib/Analysis/Convex/EGauge.lean +++ b/Mathlib/Analysis/Convex/EGauge.lean @@ -111,8 +111,7 @@ lemma egauge_zero_right (hs : s.Nonempty) : egauge 𝕜 s 0 = 0 := by have : 0 ∈ (0 : 𝕜) • s := by simp [zero_smul_set hs] simpa using egauge_le_of_mem_smul this -@[simp] -lemma egauge_zero_zero : egauge 𝕜 (0 : Set E) 0 = 0 := egauge_zero_right _ ⟨0, rfl⟩ +lemma egauge_zero_zero : egauge 𝕜 (0 : Set E) 0 = 0 := by simp lemma egauge_le_one (h : x ∈ s) : egauge 𝕜 s x ≤ 1 := by rw [← one_smul 𝕜 s] at h diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index b18e30f62e781..564cc2da28dcb 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -1044,7 +1044,7 @@ def powCardSubgroup {G : Type*} [Group G] [Fintype G] (S : Set G) (hS : S.Nonemp have one_mem : (1 : G) ∈ S ^ Fintype.card G := by obtain ⟨a, ha⟩ := hS rw [← pow_card_eq_one] - exact Set.pow_mem_pow ha (Fintype.card G) + exact Set.pow_mem_pow ha subgroupOfIdempotent (S ^ Fintype.card G) ⟨1, one_mem⟩ <| by classical apply (Set.eq_of_subset_of_card_le (Set.subset_mul_left _ one_mem) (ge_of_eq _)).symm From e6b14d47d78f3846fbf3c5d50bb1eccd71abc6e5 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Tue, 26 Nov 2024 14:42:14 +0000 Subject: [PATCH 322/829] doc(GroupTheory/Commensurable): clarify some docstrings (#19127) --- Mathlib/GroupTheory/Commensurable.lean | 16 ++++++++++------ Mathlib/GroupTheory/Index.lean | 14 +++++++------- 2 files changed, 17 insertions(+), 13 deletions(-) diff --git a/Mathlib/GroupTheory/Commensurable.lean b/Mathlib/GroupTheory/Commensurable.lean index b084767be576a..e69c86c8a8949 100644 --- a/Mathlib/GroupTheory/Commensurable.lean +++ b/Mathlib/GroupTheory/Commensurable.lean @@ -8,19 +8,23 @@ import Mathlib.GroupTheory.Index /-! # Commensurability for subgroups -This file defines commensurability for subgroups of a group `G`. It then goes on to prove that -commensurability defines an equivalence relation and finally defines the commensurator of a subgroup -of `G`. +Two subgroups `H` and `K` of a group `G` are commensurable if `H ∩ K` has finite index in both `H` +and `K`. + +This file defines commensurability for subgroups of a group `G`. It goes on to prove that +commensurability defines an equivalence relation on subgroups of `G` and finally defines the +commensurator of a subgroup `H` of `G`, which is the elements `g` of `G` such that `gHg⁻¹` is +commensurable with `H`. ## Main definitions -* `Commensurable`: defines commensurability for two subgroups `H`, `K` of `G` -* `commensurator`: defines the commensurator of a subgroup `H` of `G`. +* `Commensurable H K`: the statement that the subgroups `H` and `K` of `G` are commensurable. +* `commensurator H`: the commensurator of a subgroup `H` of `G`. ## Implementation details We define the commensurator of a subgroup `H` of `G` by first defining it as a subgroup of -`(conjAct G)`, which we call commensurator' and then taking the pre-image under +`(conjAct G)`, which we call `commensurator'` and then taking the pre-image under the map `G → (conjAct G)` to obtain our commensurator as a subgroup of `G`. -/ diff --git a/Mathlib/GroupTheory/Index.lean b/Mathlib/GroupTheory/Index.lean index c31a41f62822f..0201d5ad30137 100644 --- a/Mathlib/GroupTheory/Index.lean +++ b/Mathlib/GroupTheory/Index.lean @@ -42,16 +42,16 @@ open Cardinal variable {G G' : Type*} [Group G] [Group G'] (H K L : Subgroup G) -/-- The index of a subgroup as a natural number, and returns 0 if the index is infinite. -/ -@[to_additive "The index of a subgroup as a natural number, -and returns 0 if the index is infinite."] +/-- The index of a subgroup as a natural number. Returns `0` if the index is infinite. -/ +@[to_additive "The index of an additive subgroup as a natural number. +Returns 0 if the index is infinite."] noncomputable def index : ℕ := Nat.card (G ⧸ H) -/-- The relative index of a subgroup as a natural number, - and returns 0 if the relative index is infinite. -/ -@[to_additive "The relative index of a subgroup as a natural number, -and returns 0 if the relative index is infinite."] +/-- If `H` and `K` are subgroups of a group `G`, then `relindex H K : ℕ` is the index +of `H ∩ K` in `K`. The function returns `0` if the index is infinite. -/ +@[to_additive "If `H` and `K` are subgroups of an additive group `G`, then `relindex H K : ℕ` +is the index of `H ∩ K` in `K`. The function returns `0` if the index is infinite."] noncomputable def relindex : ℕ := (H.subgroupOf K).index From 19094c4d59858e7ff3ef353faebf2e5bd009dacb Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Tue, 26 Nov 2024 17:39:51 +0000 Subject: [PATCH 323/829] feat: handle scalar multiplication in `linear_combination` (#19136) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR extends `linear_combination` to handle scalar multiplication. Examples (under sensible typeclasses, here omitted): ```lean example {a b : K} {x : V} (h : a ^ 2 + b ^ 2 = 1) : a • (a • x - b • y) + (b • a • y + b • b • x) = x := by linear_combination (norm := module) h • x example {v w x y : V} (h₁ : x - y = -(v - w)) (h₂ : x + y = v + w) : x = w := by linear_combination (norm := module) (2:K)⁻¹ • (h₁ + h₂) example {a b : K} {x y z : V} (hb : 0 ≤ b) (hab : a + b = 1) (H : z ≤ a • x + b • y) (hyz : y ≤ z) : a • z ≤ a • x := by linear_combination (norm := skip) b • hyz + hab • z + H apply le_of_eq module ``` As is apparent from these examples, it is still generally necessary to invoke a nonstandard discharger (generally `module` or `match_scalars`) when using `linear_combination` with this new feature, since the default discharger (`ring`) is not usually useful on module goals. --- Mathlib/Analysis/Convex/Join.lean | 2 +- Mathlib/Analysis/Convex/Side.lean | 2 +- .../Analysis/InnerProductSpace/OfNorm.lean | 4 +- Mathlib/Tactic/LinearCombination.lean | 18 ++++-- Mathlib/Tactic/LinearCombination/Lemmas.lean | 56 +++++++++++++++++++ MathlibTest/linear_combination.lean | 53 +++++++++++++++++- MathlibTest/module.lean | 41 ++++---------- 7 files changed, 135 insertions(+), 41 deletions(-) diff --git a/Mathlib/Analysis/Convex/Join.lean b/Mathlib/Analysis/Convex/Join.lean index 058cec43d5537..76080ce818ed6 100644 --- a/Mathlib/Analysis/Convex/Join.lean +++ b/Mathlib/Analysis/Convex/Join.lean @@ -112,7 +112,7 @@ 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₁, ?_⟩ - linear_combination (norm := module) congr(-$hab₂ • (a₁ • x + b₁ • y)) + linear_combination (norm := module) -hab₂ • (a₁ • x + b₁ • y) refine ⟨x, hx, (a₂ * b₁ / (a₂ * b₁ + b₂)) • y + (b₂ / (a₂ * b₁ + b₂)) • z, ⟨y, hy, z, hz, _, _, by positivity, by positivity, by field_simp, rfl⟩, diff --git a/Mathlib/Analysis/Convex/Side.lean b/Mathlib/Analysis/Convex/Side.lean index 9e7135f3517d0..e623f6f31e49d 100644 --- a/Mathlib/Analysis/Convex/Side.lean +++ b/Mathlib/Analysis/Convex/Side.lean @@ -580,7 +580,7 @@ theorem wOppSide_iff_exists_wbtw {s : AffineSubspace R P} {x y : P} : · have : (r₂ / (r₁ + r₂)) • (y -ᵥ p₂ + (p₂ -ᵥ p₁) - (x -ᵥ p₁)) + (x -ᵥ p₁) = (r₂ / (r₁ + r₂)) • (p₂ -ᵥ p₁) := by rw [← neg_vsub_eq_vsub_rev p₂ y] - linear_combination (norm := match_scalars <;> field_simp) congr((r₁ + r₂)⁻¹ • $h) + linear_combination (norm := match_scalars <;> field_simp) (r₁ + r₂)⁻¹ • h 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₁ diff --git a/Mathlib/Analysis/InnerProductSpace/OfNorm.lean b/Mathlib/Analysis/InnerProductSpace/OfNorm.lean index b62522a0e45d6..374ce175c377f 100644 --- a/Mathlib/Analysis/InnerProductSpace/OfNorm.lean +++ b/Mathlib/Analysis/InnerProductSpace/OfNorm.lean @@ -121,10 +121,10 @@ theorem inner_.conj_symm (x y : E) : conj (inner_ 𝕜 y x) = inner_ 𝕜 x y := 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 convert I_smul ((I : 𝕜) • x + y) using 2 - linear_combination (norm := module) congr(-$hI' • x) + linear_combination (norm := module) -hI' • x have h₂ : ‖(I : 𝕜) • y + x‖ = ‖(I : 𝕜) • x - y‖ := by convert (I_smul ((I : 𝕜) • y + x)).symm using 2 - linear_combination (norm := module) congr(-$hI' • y) + linear_combination (norm := module) -hI' • y rw [h₁, h₂] ring diff --git a/Mathlib/Tactic/LinearCombination.lean b/Mathlib/Tactic/LinearCombination.lean index 5c891e021eae9..d5f5a2b863b2e 100644 --- a/Mathlib/Tactic/LinearCombination.lean +++ b/Mathlib/Tactic/LinearCombination.lean @@ -46,9 +46,9 @@ inductive Expanded /-- A value, equivalently a proof of `c = c`. -/ | const (c : Syntax.Term) -/-- The handling in `linear_combination` of left- and right-multiplication and of division all three -proceed according to the same logic, specified here: given a proof `p` of an (in)equality and a -constant `c`, +/-- The handling in `linear_combination` of left- and right-multiplication and scalar-multiplication +and of division all five proceed according to the same logic, specified here: given a proof `p` of +an (in)equality and a constant `c`, * if `p` is a proof of an equation, multiply/divide through by `c`; * if `p` is a proof of a non-strict inequality, run `positivity` to find a proof that `c` is nonnegative, then multiply/divide through by `c`, invoking the nonnegativity of `c` where needed; @@ -58,7 +58,7 @@ constant `c`, This generic logic takes as a parameter the object `lems`: the four lemmas corresponding to the four cases. -/ -def rescale (lems : Ineq.WithStrictness → Name) (ty : Expr) (p c : Term) : +def rescale (lems : Ineq.WithStrictness → Name) (ty : Option Expr) (p c : Term) : Ineq → TermElabM Expanded | eq => do let i := mkIdent <| lems .eq @@ -86,7 +86,8 @@ using `+`/`-`/`*`/`/` on equations and values. inequality. * `.const c` means that the input expression is not an equation but a value. -/ -partial def expandLinearCombo (ty : Expr) (stx : Syntax.Term) : TermElabM Expanded := withRef stx do +partial def expandLinearCombo (ty : Option Expr) (stx : Syntax.Term) : + TermElabM Expanded := withRef stx do match stx with | `(($e)) => expandLinearCombo ty e | `($e₁ + $e₂) => do @@ -128,6 +129,13 @@ partial def expandLinearCombo (ty : Expr) (stx : Syntax.Term) : TermElabM Expand | .const c₁, .proof rel₂ p₂ => rescale mulConstRelData ty p₂ c₁ rel₂ | .proof _ _, .proof _ _ => throwErrorAt tk "'linear_combination' supports only linear operations" + | `($e₁ •%$tk $e₂) => do + match ← expandLinearCombo none e₁, ← expandLinearCombo ty e₂ with + | .const c₁, .const c₂ => .const <$> ``($c₁ • $c₂) + | .proof rel₁ p₁, .const c₂ => rescale smulRelConstData ty p₁ c₂ rel₁ + | .const c₁, .proof rel₂ p₂ => rescale smulConstRelData none p₂ c₁ rel₂ + | .proof _ _, .proof _ _ => + throwErrorAt tk "'linear_combination' supports only linear operations" | `($e₁ /%$tk $e₂) => do match ← expandLinearCombo ty e₁, ← expandLinearCombo ty e₂ with | .const c₁, .const c₂ => .const <$> ``($c₁ / $c₂) diff --git a/Mathlib/Tactic/LinearCombination/Lemmas.lean b/Mathlib/Tactic/LinearCombination/Lemmas.lean index e22bec542ae8b..b8b2a3acda4f9 100644 --- a/Mathlib/Tactic/LinearCombination/Lemmas.lean +++ b/Mathlib/Tactic/LinearCombination/Lemmas.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Abby J. Goldberg, Mario Carneiro, Heather Macbeth -/ import Mathlib.Algebra.Order.Field.Defs +import Mathlib.Algebra.Order.Module.OrderedSMul import Mathlib.Data.Ineq /-! @@ -17,6 +18,7 @@ open Lean namespace Mathlib.Tactic.LinearCombination variable {α : Type*} {a a' a₁ a₂ b b' b₁ b₂ c : α} +variable {K : Type*} {t s : K} /-! ### Addition -/ @@ -68,6 +70,42 @@ theorem mul_const_lt_weak [OrderedSemiring α] (p : b < c) {a : α} (ha : 0 ≤ a * b ≤ a * c := mul_le_mul_of_nonneg_left p.le ha +/-! ### Scalar multiplication -/ + +theorem smul_eq_const [SMul K α] (p : t = s) (c : α) : t • c = s • c := p ▸ rfl + +theorem smul_le_const [OrderedRing K] [OrderedAddCommGroup α] [Module K α] + [OrderedSMul K α] (p : t ≤ s) {a : α} (ha : 0 ≤ a) : + t • a ≤ s • a := + smul_le_smul_of_nonneg_right p ha + +theorem smul_lt_const [OrderedRing K] [OrderedAddCommGroup α] [Module K α] + [OrderedSMul K α] (p : t < s) {a : α} (ha : 0 < a) : + t • a < s • a := + smul_lt_smul_of_pos_right p ha + +theorem smul_lt_const_weak [OrderedRing K] [OrderedAddCommGroup α] [Module K α] + [OrderedSMul K α] (p : t < s) {a : α} (ha : 0 ≤ a) : + t • a ≤ s • a := + smul_le_smul_of_nonneg_right p.le ha + +theorem smul_const_eq [SMul K α] (p : b = c) (s : K) : s • b = s • c := p ▸ rfl + +theorem smul_const_le [OrderedSemiring K] [OrderedAddCommMonoid α] [Module K α] + [OrderedSMul K α] (p : b ≤ c) {s : K} (hs : 0 ≤ s) : + s • b ≤ s • c := + smul_le_smul_of_nonneg_left p hs + +theorem smul_const_lt [OrderedSemiring K] [OrderedAddCommMonoid α] [Module K α] + [OrderedSMul K α] (p : b < c) {s : K} (hs : 0 < s) : + s • b < s • c := + smul_lt_smul_of_pos_left p hs + +theorem smul_const_lt_weak [OrderedSemiring K] [OrderedAddCommMonoid α] [Module K α] + [OrderedSMul K α] (p : b < c) {s : K} (hs : 0 ≤ s) : + s • b ≤ s • c := + smul_le_smul_of_nonneg_left p.le hs + /-! ### Division -/ theorem div_eq_const [Div α] (p : a = b) (c : α) : a / c = b / c := p ▸ rfl @@ -176,6 +214,24 @@ def mulConstRelData : Ineq.WithStrictness → Name | .lt true => ``mul_const_lt | .lt false => ``mul_const_lt_weak +/-- Given an (in)equality, look up the lemma to left-scalar-multiply it by a constant (scalar). +If relevant, also take into account the degree of positivity which can be proved of the constant: +strict or non-strict. -/ +def smulRelConstData : Ineq.WithStrictness → Name + | .eq => ``smul_eq_const + | .le => ``smul_le_const + | .lt true => ``smul_lt_const + | .lt false => ``smul_lt_const_weak + +/-- Given an (in)equality, look up the lemma to right-scalar-multiply it by a constant (vector). +If relevant, also take into account the degree of positivity which can be proved of the constant: +strict or non-strict. -/ +def smulConstRelData : Ineq.WithStrictness → Name + | .eq => ``smul_const_eq + | .le => ``smul_const_le + | .lt true => ``smul_const_lt + | .lt false => ``smul_const_lt_weak + /-- Given an (in)equality, look up the lemma to divide it by a constant. If relevant, also take into account the degree of positivity which can be proved of the constant: strict or non-strict. -/ def divRelConstData : Ineq.WithStrictness → Name diff --git a/MathlibTest/linear_combination.lean b/MathlibTest/linear_combination.lean index 4aa730f384389..b947d8190146e 100644 --- a/MathlibTest/linear_combination.lean +++ b/MathlibTest/linear_combination.lean @@ -1,6 +1,8 @@ import Mathlib.Tactic.Abel +import Mathlib.Tactic.FieldSimp import Mathlib.Tactic.LinearCombination import Mathlib.Tactic.Linarith +import Mathlib.Tactic.Module set_option autoImplicit true @@ -123,6 +125,55 @@ example (x y z w : ℚ) (hzw : z = w) : x * z + 2 * y * z = x * w + 2 * y * w := example (x y : ℤ) (h : x = 0) : y ^ 2 * x = 0 := by linear_combination y ^ 2 * h +/-! ### Scalar multiplication -/ + +section +variable {K V : Type*} + +section +variable [AddCommGroup V] [Field K] [CharZero K] [Module K V] {a b μ ν : K} {v w x y : V} + +example (h : a ^ 2 + b ^ 2 = 1) : a • (a • x - b • y) + (b • a • y + b • b • x) = x := by + linear_combination (norm := module) h • x + +example (h1 : a • x + b • y = 0) (h2 : a • μ • x + b • ν • y = 0) : (μ - ν) • a • x = 0 := by + linear_combination (norm := module) h2 - ν • h1 + +example (h₁ : x - y = -(v - w)) (h₂ : x + y = v + w) : x = w := by + linear_combination (norm := module) (2:K)⁻¹ • h₁ + (2:K)⁻¹ • h₂ + +example (h : a + b ≠ 0) (H : a • x = b • y) : x = (b / (a + b)) • (x + y) := by + linear_combination (norm := match_scalars) (a + b)⁻¹ • H + · field_simp + ring + · ring + +end + +example [OrderedCommSemiring K] [OrderedCancelAddCommMonoid V] [Module K V] [OrderedSMul K V] + {x y r : V} (hx : x < r) (hy : y < r) {a b : K} (ha : 0 < a) (hb : 0 ≤ b) (hab : a + b = 1) : + a • x + b • y < r := by + linear_combination (norm := skip) a • hx + b • hy + hab • r + apply le_of_eq + module + +example [OrderedCommSemiring K] [OrderedCancelAddCommMonoid V] [Module K V] [OrderedSMul K V] + {x y z : V} (hyz : y ≤ z) {a b : K} (hb : 0 ≤ b) (hab : a + b = 1) (H : z ≤ a • x + b • y) : + a • z ≤ a • x := by + linear_combination (norm := skip) b • hyz + hab • z + H + apply le_of_eq + module + +example [OrderedCommRing K] [OrderedAddCommGroup V] [Module K V] [OrderedSMul K V] + {x y : V} (hx : 0 < x) (hxy : x < y) {a b c : K} (hc : 0 < c) (hac : c < a) (hab : a + b ≤ 1): + c • x + b • y < y := by + have := hx.trans hxy + linear_combination (norm := skip) hab • y + hac • y + c • hxy + apply le_of_eq + module + +end + /-! ### Tests in semirings -/ example (a _b : ℕ) (h1 : a = 3) : a = 3 := by @@ -470,7 +521,7 @@ typeclass inference is demanded by the lemmas it orchestrates. This example too (and 73 ms on a good laptop) on an implementation with "minimal" typeclasses everywhere, e.g. lots of `CovariantClass`/`ContravariantClass`, and takes 206 heartbeats (10 ms on a good laptop) on the implementation at the time of joining Mathlib (November 2024). -/ -set_option maxHeartbeats 1000 in +set_option maxHeartbeats 1100 in example {a b : ℝ} (h : a < b) : 0 < b - a := by linear_combination (norm := skip) h exact test_sorry diff --git a/MathlibTest/module.lean b/MathlibTest/module.lean index 2736fdc32d143..8c96ddc566272 100644 --- a/MathlibTest/module.lean +++ b/MathlibTest/module.lean @@ -156,12 +156,8 @@ example (h : a = b) : a • x = b • x := by match_scalars linear_combination h -/- `linear_combination` does not currently handle `•`. The following mimics what should eventually -be performed by a `linear_combination` call, with exact syntax TBD -- maybe -`linear_combination (norm := module) h • x` or `module_combination h • x`. -/ example (h : a = b) : a • x = b • x := by - apply eq_of_eq (congr($h • x):) - module + linear_combination (norm := module) h • x example (h : a ^ 2 + b ^ 2 = 1) : a • (a • x - b • y) + (b • a • y + b • b • x) = x := by match_scalars @@ -169,25 +165,17 @@ example (h : a ^ 2 + b ^ 2 = 1) : a • (a • x - b • y) + (b • a • y + b · ring example (h : a ^ 2 + b ^ 2 = 1) : a • (a • x - b • y) + (b • a • y + b • b • x) = x := by - -- `linear_combination (norm := module) h • x` - apply eq_of_eq (congr($h • x):) - module + linear_combination (norm := module) h • x example (h1 : a • x + b • y = 0) (h2 : a • μ • x + b • ν • y = 0) : (μ - ν) • a • x = 0 ∧ (μ - ν) • b • y = 0 := by constructor - · -- `linear_combination (norm := module) h2 - ν • h1` - apply eq_of_eq (congr($h2 - ν • $h1):) - module - · -- `linear_combination (norm := module) μ • h1 + h2` - apply eq_of_eq (congr(μ • $h1 - $h2):) - module + · linear_combination (norm := module) h2 - ν • h1 + · linear_combination (norm := module) μ • h1 - h2 example (h1 : 0 • z + a • x + b • y = 0) (h2 : 0 • ρ • z + a • μ • x + b • ν • y = 0) : (μ - ν) • a • x = 0 := by - -- `linear_combination (norm := module) h2 - ν • h1` - apply eq_of_eq (congr($h2 - ν • $h1):) - module + linear_combination (norm := module) h2 - ν • h1 example (h1 : a • x + b • y + c • z = 0) @@ -196,15 +184,9 @@ example (μ - ν) • (μ - ρ) • a • x = 0 ∧ (μ - ν) • (ν - ρ) • b • y = 0 ∧ (μ - ρ) • (ν - ρ) • c • z = 0 := by refine ⟨?_, ?_, ?_⟩ - · -- `linear_combination (norm := module) h3 - (ν + ρ) • h2 + ν • ρ • h1` - apply eq_of_eq (congr($h3 - (ν + ρ) • $h2 + ν • ρ • $h1):) - module - · -- `linear_combination (norm := module) - h3 + (μ + ρ) • h2 - μ • ρ • h1` - apply eq_of_eq (congr(- $h3 + (μ + ρ) • $h2 - μ • ρ • $h1):) - module - · -- `linear_combination (norm := module) h3 - (μ + ν) • h2 + μ • ν • h1` - apply eq_of_eq (congr($h3 - (μ + ν) • $h2 + μ • ν • $h1):) - module + · linear_combination (norm := module) h3 - (ν + ρ) • h2 + ν • ρ • h1 + · linear_combination (norm := module) - h3 + (μ + ρ) • h2 - μ • ρ • h1 + · linear_combination (norm := module) h3 - (μ + ν) • h2 + μ • ν • h1 /-- error: ring failed, ring expressions not equal @@ -253,9 +235,7 @@ variable [Field K] [CharZero K] [Module K V] example : (2:K)⁻¹ • x + (3:K)⁻¹ • x + (6:K)⁻¹ • x = x := by module example (h₁ : t - u = -(v - w)) (h₂ : t + u = v + w) : t = w := by - -- `linear_combination (norm := module) 2⁻¹ • h₁ + 2⁻¹ • h₂` - apply eq_of_eq (congr((2:K)⁻¹ • $h₁ + (2:K)⁻¹ • $h₂):) - module + linear_combination (norm := module) (2:K)⁻¹ • h₁ + (2:K)⁻¹ • h₂ end CharZeroField @@ -288,8 +268,7 @@ example (h₁ : 1 = a ^ 2 + b ^ 2) (h₂ : 1 - a ≠ 0) : ((2 / (1 - a)) ^ 2 * b ^ 2 + 4)⁻¹ • (4:K) • ((2 / (1 - a)) • y) + ((2 / (1 - a)) ^ 2 * b ^ 2 + 4)⁻¹ • ((2 / (1 - a)) ^ 2 * b ^ 2 - 4) • x = a • x + y := by - -- `linear_combination (norm := skip) (h₁ * (b ^ 2 + (1 - a) ^ 2)⁻¹) • (y + (a - 1) • x)` - apply eq_of_eq (congr(($h₁ * (b ^ 2 + (1 - a) ^ 2)⁻¹) • (y + (a - 1) • x)):) + linear_combination (norm := skip) (h₁ * (b ^ 2 + (1 - a) ^ 2)⁻¹) • (y + (a - 1) • x) match_scalars · field_simp ring From 38d4438a8695ba746705bb18a58eb1d104febfb8 Mon Sep 17 00:00:00 2001 From: Nick Ward <102917377+gio256@users.noreply.github.com> Date: Tue, 26 Nov 2024 19:59:18 +0000 Subject: [PATCH 324/829] chore(AlgebraicTopology): move Nerve.* to CategoryTheory.Nerve.* (#19521) Moves definitions to more appropriate namespaces: - Nerve.strictSegal -> CategoryTheory.Nerve.strictSegal - Nerve.quasicategory -> CategoryTheory.Nerve.quasicategory Co-Authored-By: [Emily Riehl](https://github.com/emilyriehl) --- Mathlib/AlgebraicTopology/Quasicategory/Nerve.lean | 6 +++--- Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean | 6 ++---- 2 files changed, 5 insertions(+), 7 deletions(-) diff --git a/Mathlib/AlgebraicTopology/Quasicategory/Nerve.lean b/Mathlib/AlgebraicTopology/Quasicategory/Nerve.lean index 22963dfe7b25f..5b8e8e517bd47 100644 --- a/Mathlib/AlgebraicTopology/Quasicategory/Nerve.lean +++ b/Mathlib/AlgebraicTopology/Quasicategory/Nerve.lean @@ -19,13 +19,13 @@ a quasicategory. universe v u -open CategoryTheory SSet +open SSet -namespace Nerve +namespace CategoryTheory.Nerve /-- By virtue of satisfying the `StrictSegal` condition, the nerve of a category is a `Quasicategory`. -/ instance quasicategory {C : Type u} [Category.{v} C] : Quasicategory (nerve C) := inferInstance -end Nerve +end CategoryTheory.Nerve diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean b/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean index 71e41b710ae25..7436832a79255 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean @@ -156,12 +156,10 @@ end StrictSegal end SSet -namespace Nerve +namespace CategoryTheory.Nerve open SSet -variable {C : Type*} [Category C] {n : ℕ} - /-- Simplices in the nerve of categories are uniquely determined by their spine. Indeed, this property describes the essential image of the nerve functor.-/ noncomputable instance strictSegal (C : Type u) [Category.{v} C] : StrictSegal (nerve C) where @@ -184,4 +182,4 @@ noncomputable instance strictSegal (C : Type u) [Category.{v} C] : StrictSegal ( · intro i hi apply ComposableArrows.mkOfObjOfMapSucc_map_succ -end Nerve +end CategoryTheory.Nerve From b25c29685772c4e73f377e6746ed44c2d3d29907 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Tue, 26 Nov 2024 20:08:00 +0000 Subject: [PATCH 325/829] chore(RingTheory/Localization): golfs and generalizations (#19310) + In RingTheory/LocalProperties/Basic: add some lemmas connecting localization of submodules and localization of ideals, and use results about the former to golf results about the latter. + In RingTheory/Localization/Module: golf a proof using the recently introduced IsLocalization.linearMap_compatibleSMul. + In Algebra/Module/Projective, Algebra/Module/LocalizedModule/IsLocalization, RingTheory/Localization/Submodule: generalize from AddCommGroup/CommRing to AddCommMonoid/CommSemiring. Also golf a lemma statement in the second file. + RingTheory/Localization/Algebra, RingTheory/Localization/Ideal: add two TODOs that are not done here for import reasons. --- .../LocalizedModule/IsLocalization.lean | 4 +- Mathlib/Algebra/Module/Projective.lean | 8 +- Mathlib/RingTheory/LocalProperties/Basic.lean | 82 ++++++++++--------- Mathlib/RingTheory/Localization/Algebra.lean | 1 + Mathlib/RingTheory/Localization/Ideal.lean | 1 + Mathlib/RingTheory/Localization/Module.lean | 17 +--- .../RingTheory/Localization/Submodule.lean | 20 +++-- 7 files changed, 65 insertions(+), 68 deletions(-) diff --git a/Mathlib/Algebra/Module/LocalizedModule/IsLocalization.lean b/Mathlib/Algebra/Module/LocalizedModule/IsLocalization.lean index fb79c4a5368da..aed527565f88b 100644 --- a/Mathlib/Algebra/Module/LocalizedModule/IsLocalization.lean +++ b/Mathlib/Algebra/Module/LocalizedModule/IsLocalization.lean @@ -34,11 +34,11 @@ instance {A Aₛ} [CommSemiring A] [Algebra R A][CommSemiring Aₛ] [Algebra A A isLocalizedModule_iff_isLocalization.mpr h lemma isLocalizedModule_iff_isLocalization' (R') [CommSemiring R'] [Algebra R R'] : - IsLocalizedModule S (Algebra.ofId R R').toLinearMap ↔ IsLocalization S R' := by + IsLocalizedModule S (Algebra.linearMap R R') ↔ IsLocalization S R' := by convert isLocalizedModule_iff_isLocalization (S := S) (A := R) (Aₛ := R') exact (Submonoid.map_id S).symm -instance {A} [CommRing A] [Algebra R A] [IsLocalization S A] : +instance {A} [CommSemiring A] [Algebra R A] [IsLocalization S A] : IsLocalizedModule S (Algebra.linearMap R A) := (isLocalizedModule_iff_isLocalization' S _).mpr inferInstance diff --git a/Mathlib/Algebra/Module/Projective.lean b/Mathlib/Algebra/Module/Projective.lean index fdce19a9a6937..42ac6dfc4685e 100644 --- a/Mathlib/Algebra/Module/Projective.lean +++ b/Mathlib/Algebra/Module/Projective.lean @@ -194,13 +194,13 @@ end Semiring section Ring -variable {R : Type u} [Ring R] {P : Type v} [AddCommMonoid P] [Module R P] -variable {R₀ M N} [CommRing R₀] [Algebra R₀ R] [AddCommGroup M] [Module R₀ M] [Module R M] -variable [IsScalarTower R₀ R M] [AddCommGroup N] [Module R₀ N] +variable {R : Type u} [Semiring R] {P : Type v} [AddCommMonoid P] [Module R P] +variable {R₀ M N} [CommSemiring R₀] [Algebra R₀ R] [AddCommMonoid M] [Module R₀ M] [Module R M] +variable [IsScalarTower R₀ R M] [AddCommMonoid N] [Module R₀ N] /-- A module is projective iff it is the direct summand of a free module. -/ theorem Projective.iff_split : Module.Projective R P ↔ - ∃ (M : Type max u v) (_ : AddCommGroup M) (_ : Module R M) (_ : Module.Free R M) + ∃ (M : Type max u v) (_ : AddCommMonoid M) (_ : Module R M) (_ : Module.Free R M) (i : P →ₗ[R] M) (s : M →ₗ[R] P), s.comp i = LinearMap.id := ⟨fun ⟨i, hi⟩ ↦ ⟨P →₀ R, _, _, inferInstance, i, Finsupp.linearCombination R id, LinearMap.ext hi⟩, fun ⟨_, _, _, _, i, s, H⟩ ↦ Projective.of_split i s H⟩ diff --git a/Mathlib/RingTheory/LocalProperties/Basic.lean b/Mathlib/RingTheory/LocalProperties/Basic.lean index 0b332d9f25b8b..354847e24b104 100644 --- a/Mathlib/RingTheory/LocalProperties/Basic.lean +++ b/Mathlib/RingTheory/LocalProperties/Basic.lean @@ -3,10 +3,10 @@ Copyright (c) 2021 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ -import Mathlib.RingTheory.Ideal.Colon import Mathlib.RingTheory.Localization.AtPrime import Mathlib.RingTheory.Localization.BaseChange import Mathlib.RingTheory.Localization.Submodule +import Mathlib.RingTheory.LocalProperties.Submodule import Mathlib.RingTheory.RingHomProperties /-! @@ -45,12 +45,12 @@ open scoped Pointwise Classical universe u +section Properties + variable {R S : Type u} [CommRing R] [CommRing S] (f : R →+* S) variable (R' S' : Type u) [CommRing R'] [CommRing S'] variable [Algebra R R'] [Algebra S S'] -section Properties - section CommRing variable (P : ∀ (R : Type u) [CommRing R], Prop) @@ -396,7 +396,35 @@ end Properties section Ideal -open scoped nonZeroDivisors +variable {R : Type*} (S : Type*) [CommSemiring R] [CommSemiring S] [Algebra R S] +variable (p : Submonoid R) [IsLocalization p S] + +theorem Ideal.localized'_eq_map (I : Ideal R) : + haveI := (isLocalizedModule_iff_isLocalization' p S).mpr inferInstance + Submodule.localized' S p (Algebra.linearMap R S) I = I.map (algebraMap R S) := + SetLike.ext fun x ↦ by + simp_rw [Submodule.mem_localized', IsLocalization.mem_map_algebraMap_iff p, + IsLocalizedModule.mk'_eq_iff, mul_comm x, eq_comm (a := _ * x), ← Algebra.smul_def, + Prod.exists, Subtype.exists, ← exists_prop] + rfl + +theorem Ideal.localized₀_eq_restrictScalars_map (I : Ideal R) : + Submodule.localized₀ p (Algebra.linearMap R S) I = (I.map (algebraMap R S)).restrictScalars R := + congr(Submodule.restrictScalars R $(localized'_eq_map S p I)) + +theorem Algebra.idealMap_eq_ofEq_comp_toLocalized₀ (I : Ideal R) : + haveI := (isLocalizedModule_iff_isLocalization' p S).mpr inferInstance + Algebra.idealMap S I = + (LinearEquiv.ofEq _ _ <| Ideal.localized₀_eq_restrictScalars_map S p I).toLinearMap ∘ₗ + Submodule.toLocalized₀ p (Algebra.linearMap R S) I := + rfl + +theorem Ideal.mem_of_localization_maximal {r : R} {J : Ideal R} + (h : ∀ (P : Ideal R) (_ : P.IsMaximal), + algebraMap R _ r ∈ Ideal.map (algebraMap R (Localization.AtPrime P)) J) : + r ∈ J := + Submodule.mem_of_localization_maximal _ _ _ _ fun P hP ↦ by + apply (localized'_eq_map (Localization.AtPrime P) P.primeCompl J).symm ▸ h P hP /-- Let `I J : Ideal R`. If the localization of `I` at each maximal ideal `P` is included in the localization of `J` at `P`, then `I ≤ J`. -/ @@ -404,22 +432,8 @@ theorem Ideal.le_of_localization_maximal {I J : Ideal R} (h : ∀ (P : Ideal R) (_ : P.IsMaximal), Ideal.map (algebraMap R (Localization.AtPrime P)) I ≤ Ideal.map (algebraMap R (Localization.AtPrime P)) J) : - I ≤ J := by - intro x hx - suffices J.colon (Ideal.span {x}) = ⊤ by - simpa using Submodule.mem_colon.mp - (show (1 : R) ∈ J.colon (Ideal.span {x}) from this.symm ▸ Submodule.mem_top) x - (Ideal.mem_span_singleton_self x) - refine Not.imp_symm (J.colon (Ideal.span {x})).exists_le_maximal ?_ - push_neg - intro P hP le - obtain ⟨⟨⟨a, ha⟩, ⟨s, hs⟩⟩, eq⟩ := - (IsLocalization.mem_map_algebraMap_iff P.primeCompl _).mp (h P hP (Ideal.mem_map_of_mem _ hx)) - rw [← _root_.map_mul, ← sub_eq_zero, ← map_sub] at eq - obtain ⟨⟨m, hm⟩, eq⟩ := (IsLocalization.map_eq_zero_iff P.primeCompl _ _).mp eq - refine hs ((hP.isPrime.mem_or_mem (le (Ideal.mem_colon_singleton.mpr ?_))).resolve_right hm) - simp only [Subtype.coe_mk, mul_sub, sub_eq_zero, mul_comm x s, mul_left_comm] at eq - simpa only [mul_assoc, eq] using J.mul_mem_left m ha + I ≤ J := + fun _ hm ↦ mem_of_localization_maximal fun P hP ↦ h P hP (mem_map_of_mem _ hm) /-- Let `I J : Ideal R`. If the localization of `I` at each maximal ideal `P` is equal to the localization of `J` at `P`, then `I = J`. -/ @@ -428,8 +442,8 @@ theorem Ideal.eq_of_localization_maximal {I J : Ideal R} Ideal.map (algebraMap R (Localization.AtPrime P)) I = Ideal.map (algebraMap R (Localization.AtPrime P)) J) : I = J := - le_antisymm (Ideal.le_of_localization_maximal fun P hP => (h P hP).le) - (Ideal.le_of_localization_maximal fun P hP => (h P hP).ge) + le_antisymm (le_of_localization_maximal fun P hP ↦ (h P hP).le) + (le_of_localization_maximal fun P hP ↦ (h P hP).ge) /-- An ideal is trivial if its localization at every maximal ideal is trivial. -/ theorem ideal_eq_bot_of_localization' (I : Ideal R) @@ -438,28 +452,16 @@ theorem ideal_eq_bot_of_localization' (I : Ideal R) I = ⊥ := Ideal.eq_of_localization_maximal fun P hP => by simpa using h P hP --- TODO: This proof should work for all modules, once we have enough material on submodules of --- localized modules. +theorem eq_zero_of_localization (r : R) + (h : ∀ (J : Ideal R) (_ : J.IsMaximal), algebraMap R (Localization.AtPrime J) r = 0) : + r = 0 := + Module.eq_zero_of_localization_maximal _ (fun _ _ ↦ Algebra.linearMap R _) r h + /-- An ideal is trivial if its localization at every maximal ideal is trivial. -/ theorem ideal_eq_bot_of_localization (I : Ideal R) (h : ∀ (J : Ideal R) (_ : J.IsMaximal), IsLocalization.coeSubmodule (Localization.AtPrime J) I = ⊥) : I = ⊥ := - ideal_eq_bot_of_localization' _ fun P hP => - (Ideal.map_eq_bot_iff_le_ker _).mpr fun x hx => by - rw [RingHom.mem_ker, ← Submodule.mem_bot R, ← h P hP, IsLocalization.mem_coeSubmodule] - exact ⟨x, hx, rfl⟩ - -theorem eq_zero_of_localization (r : R) - (h : ∀ (J : Ideal R) (_ : J.IsMaximal), algebraMap R (Localization.AtPrime J) r = 0) : - r = 0 := by - rw [← Ideal.span_singleton_eq_bot] - apply ideal_eq_bot_of_localization - intro J hJ - delta IsLocalization.coeSubmodule - erw [Submodule.map_span, Submodule.span_eq_bot] - rintro _ ⟨_, h', rfl⟩ - cases Set.mem_singleton_iff.mpr h' - exact h J hJ + bot_unique fun r hr ↦ eq_zero_of_localization r fun J hJ ↦ (h J hJ).le ⟨r, hr, rfl⟩ end Ideal diff --git a/Mathlib/RingTheory/Localization/Algebra.lean b/Mathlib/RingTheory/Localization/Algebra.lean index 0210fb6bc2be8..b52b401bdd5bf 100644 --- a/Mathlib/RingTheory/Localization/Algebra.lean +++ b/Mathlib/RingTheory/Localization/Algebra.lean @@ -32,6 +32,7 @@ variable {R S P : Type*} (Q : Type*) [CommSemiring R] [CommSemiring S] [CommSemi open IsLocalization in variable (M S) in /-- The span of `I` in a localization of `R` at `M` is the localization of `I` at `M`. -/ +-- TODO: golf using `Ideal.localized'_eq_map` instance Algebra.idealMap_isLocalizedModule (I : Ideal R) : IsLocalizedModule M (Algebra.idealMap I (S := S)) where map_units x := diff --git a/Mathlib/RingTheory/Localization/Ideal.lean b/Mathlib/RingTheory/Localization/Ideal.lean index 4d5fd06147aed..126d2cbf0c729 100644 --- a/Mathlib/RingTheory/Localization/Ideal.lean +++ b/Mathlib/RingTheory/Localization/Ideal.lean @@ -28,6 +28,7 @@ variable [Algebra R S] [IsLocalization M S] In practice, this ideal differs only in that the carrier set is defined explicitly. This definition is only meant to be used in proving `mem_map_algebraMap_iff`, and any proof that needs to refer to the explicit carrier set should use that theorem. -/ +-- TODO: golf this using `Submodule.localized'` private def map_ideal (I : Ideal R) : Ideal S where carrier := { z : S | ∃ x : I × M, z * algebraMap R S x.2 = algebraMap R S x.1 } zero_mem' := ⟨⟨0, 1⟩, by simp⟩ diff --git a/Mathlib/RingTheory/Localization/Module.lean b/Mathlib/RingTheory/Localization/Module.lean index 328a15ccb228b..40c9dccaad408 100644 --- a/Mathlib/RingTheory/Localization/Module.lean +++ b/Mathlib/RingTheory/Localization/Module.lean @@ -207,17 +207,7 @@ open IsLocalization def LinearMap.extendScalarsOfIsLocalization (f : M →ₗ[R] N) : M →ₗ[A] N where toFun := f map_add' := f.map_add - map_smul' := by - intro r m - simp only [RingHom.id_apply] - rcases mk'_surjective S r with ⟨r, s, rfl⟩ - calc f (mk' A r s • m) - = ((s : R) • mk' A 1 s) • f (mk' A r s • m) := by simp - _ = (mk' A 1 s) • (s : R) • f (mk' A r s • m) := by rw [smul_comm, smul_assoc] - _ = (mk' A 1 s) • f ((s : R) • mk' A r s • m) := by simp - _ = (mk' A 1 s) • f (r • m) := by rw [← smul_assoc, smul_mk'_self, algebraMap_smul] - _ = (mk' A 1 s) • r • f m := by simp - _ = mk' A r s • f m := by rw [smul_comm, ← smul_assoc, smul_mk'_one] + map_smul' := (IsLocalization.linearMap_compatibleSMul S A M N).map_smul _ @[simp] lemma LinearMap.restrictScalars_extendScalarsOfIsLocalization (f : M →ₗ[R] N) : (f.extendScalarsOfIsLocalization S A).restrictScalars R = f := rfl @@ -257,8 +247,7 @@ variable [IsScalarTower R Rₛ M'] [IsScalarTower R Rₛ N'] [IsLocalization S R @[simps!] noncomputable def mapExtendScalars : (M →ₗ[R] N) →ₗ[R] (M' →ₗ[Rₛ] N') := - ((LinearMap.extendScalarsOfIsLocalizationEquiv - S Rₛ).restrictScalars R).toLinearMap.comp (map S f g) + ((LinearMap.extendScalarsOfIsLocalizationEquiv S Rₛ).restrictScalars R).toLinearMap ∘ₗ map S f g end IsLocalizedModule @@ -273,7 +262,7 @@ noncomputable def LocalizedModule.map : (M →ₗ[R] N) →ₗ[R] (LocalizedModule S M →ₗ[Localization S] LocalizedModule S N) := IsLocalizedModule.mapExtendScalars S (LocalizedModule.mkLinearMap S M) - (LocalizedModule.mkLinearMap S N) (Localization S) + (LocalizedModule.mkLinearMap S N) (Localization S) @[simp] lemma LocalizedModule.map_mk (f : M →ₗ[R] N) (x y) : diff --git a/Mathlib/RingTheory/Localization/Submodule.lean b/Mathlib/RingTheory/Localization/Submodule.lean index ec8df59b713d2..a1e3a22a8e03c 100644 --- a/Mathlib/RingTheory/Localization/Submodule.lean +++ b/Mathlib/RingTheory/Localization/Submodule.lean @@ -20,7 +20,7 @@ commutative ring, field of fractions -/ -variable {R : Type*} [CommRing R] (M : Submonoid R) (S : Type*) [CommRing S] +variable {R : Type*} [CommSemiring R] (M : Submonoid R) (S : Type*) [CommSemiring S] variable [Algebra R S] namespace IsLocalization @@ -59,7 +59,7 @@ theorem coeSubmodule_mul (I J : Ideal R) : theorem coeSubmodule_fg (hS : Function.Injective (algebraMap R S)) (I : Ideal R) : Submodule.FG (coeSubmodule S I) ↔ Submodule.FG I := - ⟨Submodule.fg_of_fg_map _ (LinearMap.ker_eq_bot.mpr hS), Submodule.FG.map _⟩ + ⟨Submodule.fg_of_fg_map_injective _ hS, Submodule.FG.map _⟩ @[simp] theorem coeSubmodule_span (s : Set R) : @@ -78,7 +78,10 @@ theorem isNoetherianRing (h : IsNoetherianRing R) : IsNoetherianRing S := by rw [isNoetherianRing_iff, isNoetherian_iff] at h ⊢ exact OrderEmbedding.wellFounded (IsLocalization.orderEmbedding M S).dual h -variable {S M} +section NonZeroDivisors + +variable {R : Type*} [CommRing R] {M : Submonoid R} + {S : Type*} [CommRing S] [Algebra R S] [IsLocalization M S] @[mono] theorem coeSubmodule_le_coeSubmodule (h : M ≤ nonZeroDivisors R) {I J : Ideal R} : @@ -87,7 +90,6 @@ theorem coeSubmodule_le_coeSubmodule (h : M ≤ nonZeroDivisors R) {I J : Ideal Submodule.map_le_map_iff_of_injective (f := Algebra.linearMap R S) (IsLocalization.injective _ h) _ _ - @[mono] theorem coeSubmodule_strictMono (h : M ≤ nonZeroDivisors R) : StrictMono (coeSubmodule S : Ideal R → Submodule R S) := @@ -109,9 +111,11 @@ theorem coeSubmodule_isPrincipal {I : Ideal R} (h : M ≤ nonZeroDivisors R) : · refine ⟨⟨algebraMap R S x, ?_⟩⟩ rw [hx, Ideal.submodule_span_eq, coeSubmodule_span_singleton] -variable {S} (M) +end NonZeroDivisors + +variable {S} -theorem mem_span_iff {N : Type*} [AddCommGroup N] [Module R N] [Module S N] [IsScalarTower R S N] +theorem mem_span_iff {N : Type*} [AddCommMonoid N] [Module R N] [Module S N] [IsScalarTower R S N] {x : N} {a : Set N} : x ∈ Submodule.span S a ↔ ∃ y ∈ Submodule.span R a, ∃ z : M, x = mk' S 1 z • y := by constructor @@ -155,11 +159,11 @@ namespace IsFractionRing open IsLocalization -variable {K : Type*} +variable {R K : Type*} section CommRing -variable [CommRing K] [Algebra R K] [IsFractionRing R K] +variable [CommRing R] [CommRing K] [Algebra R K] [IsFractionRing R K] @[simp, mono] theorem coeSubmodule_le_coeSubmodule {I J : Ideal R} : From c8c981908a997925fb9110b09da9f0d13a02b808 Mon Sep 17 00:00:00 2001 From: Mac Malone Date: Tue, 26 Nov 2024 22:10:33 +0000 Subject: [PATCH 326/829] fix: build cache issues (#19524) Fixes compatibility issues between Cache and Lake's builtin build cache. --- .github/build.in.yml | 9 +++++++++ .github/workflows/bors.yml | 9 +++++++++ .github/workflows/build.yml | 9 +++++++++ .github/workflows/build_fork.yml | 9 +++++++++ Cache/IO.lean | 3 ++- Cache/Requests.lean | 14 +++++++++++++- 6 files changed, 51 insertions(+), 2 deletions(-) diff --git a/.github/build.in.yml b/.github/build.in.yml index bab96d693f6b2..141b5316926ab 100644 --- a/.github/build.in.yml +++ b/.github/build.in.yml @@ -4,6 +4,14 @@ ### NB: and regenerate those files by manually running ### NB: .github/workflows/mk_build_yml.sh +env: + # Disable Lake's automatic fetching of cloud builds. + # Lake's cache is currently incompatible with Mathlib's `lake exe cache`. + # This is because Mathlib's Cache assumes all build aritfacts present in the build directory + # are valid by-products of the Mathlib build. Build artifacts fetched from Lake's cache do + # not necessarily satisfy this property. + LAKE_NO_CACHE: true + jobs: # Cancels previous runs of jobs in this file cancel: @@ -121,6 +129,7 @@ jobs: - name: prune ProofWidgets .lake run: | + lake build proofwidgets:release # The ProofWidgets release contains not just the `.js` (which we need in order to build) # but also `.oleans`, which may have been built with the wrong toolchain. # This removes them. diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index e59f48d2c7267..3bd02efbb94dd 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -14,6 +14,14 @@ on: name: continuous integration (staging) +env: + # Disable Lake's automatic fetching of cloud builds. + # Lake's cache is currently incompatible with Mathlib's `lake exe cache`. + # This is because Mathlib's Cache assumes all build aritfacts present in the build directory + # are valid by-products of the Mathlib build. Build artifacts fetched from Lake's cache do + # not necessarily satisfy this property. + LAKE_NO_CACHE: true + jobs: # Cancels previous runs of jobs in this file cancel: @@ -131,6 +139,7 @@ jobs: - name: prune ProofWidgets .lake run: | + lake build proofwidgets:release # The ProofWidgets release contains not just the `.js` (which we need in order to build) # but also `.oleans`, which may have been built with the wrong toolchain. # This removes them. diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 1c778a5720c23..b5f11a1ea1f03 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -21,6 +21,14 @@ on: name: continuous integration +env: + # Disable Lake's automatic fetching of cloud builds. + # Lake's cache is currently incompatible with Mathlib's `lake exe cache`. + # This is because Mathlib's Cache assumes all build aritfacts present in the build directory + # are valid by-products of the Mathlib build. Build artifacts fetched from Lake's cache do + # not necessarily satisfy this property. + LAKE_NO_CACHE: true + jobs: # Cancels previous runs of jobs in this file cancel: @@ -138,6 +146,7 @@ jobs: - name: prune ProofWidgets .lake run: | + lake build proofwidgets:release # The ProofWidgets release contains not just the `.js` (which we need in order to build) # but also `.oleans`, which may have been built with the wrong toolchain. # This removes them. diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index e92f87fd50a68..4da591010b8b7 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -18,6 +18,14 @@ on: name: continuous integration (mathlib forks) +env: + # Disable Lake's automatic fetching of cloud builds. + # Lake's cache is currently incompatible with Mathlib's `lake exe cache`. + # This is because Mathlib's Cache assumes all build aritfacts present in the build directory + # are valid by-products of the Mathlib build. Build artifacts fetched from Lake's cache do + # not necessarily satisfy this property. + LAKE_NO_CACHE: true + jobs: # Cancels previous runs of jobs in this file cancel: @@ -135,6 +143,7 @@ jobs: - name: prune ProofWidgets .lake run: | + lake build proofwidgets:release # The ProofWidgets release contains not just the `.js` (which we need in order to build) # but also `.oleans`, which may have been built with the wrong toolchain. # This removes them. diff --git a/Cache/IO.lean b/Cache/IO.lean index 522364f147d94..358683a890f1c 100644 --- a/Cache/IO.lean +++ b/Cache/IO.lean @@ -92,6 +92,7 @@ abbrev PackageDirs := Lean.RBMap String FilePath compare structure CacheM.Context where mathlibDepPath : FilePath packageDirs : PackageDirs + proofWidgetsBuildDir : FilePath abbrev CacheM := ReaderT CacheM.Context IO @@ -141,7 +142,7 @@ private def CacheM.getContext : IO CacheM.Context := do ("ImportGraph", LAKEPACKAGESDIR / "importGraph"), ("LeanSearchClient", LAKEPACKAGESDIR / "LeanSearchClient"), ("Plausible", LAKEPACKAGESDIR / "plausible") - ]⟩ + ], LAKEPACKAGESDIR / "proofwidgets" / ".lake" / "build"⟩ def CacheM.run (f : CacheM α) : IO α := do ReaderT.run f (← getContext) diff --git a/Cache/Requests.lean b/Cache/Requests.lean index 71b74cb4b051d..295d65cd96693 100644 --- a/Cache/Requests.lean +++ b/Cache/Requests.lean @@ -159,11 +159,23 @@ into the `lean-toolchain` file at the root directory of your project" IO.Process.exit 1 return () +/-- Fetches the ProofWidgets cloud release and prunes non-JS files. -/ +def getProofWidgets (buildDir : FilePath) : IO Unit := do + if (← buildDir.pathExists) then return + -- Unpack the ProofWidgets cloud release (for its `.js` files) + let exitCode ← (← IO.Process.spawn {cmd := "lake", args := #["-q", "build", "proofwidgets:release"]}).wait + if exitCode != 0 then + throw <| IO.userError s!"Failed to fetch ProofWidgets cloud release: lake failed with error code {exitCode}" + -- prune non-js ProofWidgets files (e.g., `olean`, `.c`) + IO.FS.removeDirAll (buildDir / "lib") + IO.FS.removeDirAll (buildDir / "ir") + /-- Downloads missing files, and unpacks files. -/ def getFiles (hashMap : IO.HashMap) (forceDownload forceUnpack parallel decompress : Bool) : IO.CacheM Unit := do let isMathlibRoot ← IO.isMathlibRoot - if !isMathlibRoot then checkForToolchainMismatch + unless isMathlibRoot do checkForToolchainMismatch + getProofWidgets (← read).proofWidgetsBuildDir downloadFiles hashMap forceDownload parallel if decompress then IO.unpackCache hashMap forceUnpack From 21af8545cb9daf04f22cddc30a99fd7dbbf3543b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Tue, 26 Nov 2024 22:19:05 +0000 Subject: [PATCH 327/829] chore: adaptation to batteries#1055 (#19487) Co-authored-by: F. G. Dorais Co-authored-by: Kim Morrison Co-authored-by: Matthew Ballard --- Mathlib/Data/Fin/VecNotation.lean | 1 - Mathlib/Data/List/Defs.lean | 8 -------- Mathlib/Data/List/FinRange.lean | 18 +----------------- Mathlib/ModelTheory/Encoding.lean | 6 +++--- lake-manifest.json | 2 +- lakefile.lean | 3 +-- 6 files changed, 6 insertions(+), 32 deletions(-) diff --git a/Mathlib/Data/Fin/VecNotation.lean b/Mathlib/Data/Fin/VecNotation.lean index 9d31131cdaf6e..1b8ca2379a123 100644 --- a/Mathlib/Data/Fin/VecNotation.lean +++ b/Mathlib/Data/Fin/VecNotation.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ import Mathlib.Data.Fin.Tuple.Basic -import Mathlib.Data.List.Defs /-! # Matrix and vector notation diff --git a/Mathlib/Data/List/Defs.lean b/Mathlib/Data/List/Defs.lean index 4b86b7218cdfb..b340e9b270a42 100644 --- a/Mathlib/Data/List/Defs.lean +++ b/Mathlib/Data/List/Defs.lean @@ -492,14 +492,6 @@ end MapAccumr set_option allowUnsafeReducibility true in attribute [semireducible] Fin.foldr.loop -/-- All elements of `Fin n`, from `0` to `n-1`. The corresponding finset is `Finset.univ`. -/ --- Note that we use `ofFn (fun x => x)` instead of `ofFn id` to avoid leaving `id` in the terms --- for e.g. `Fintype (Fin n)`. -def finRange (n : Nat) : List (Fin n) := ofFn (fun x => x) - --- Verify that `finRange` is semireducible. -example : finRange 3 = [0, 1, 2] := rfl - section Deprecated @[deprecated List.mem_cons (since := "2024-08-10")] diff --git a/Mathlib/Data/List/FinRange.lean b/Mathlib/Data/List/FinRange.lean index bf592cfe78937..83ccab9a12979 100644 --- a/Mathlib/Data/List/FinRange.lean +++ b/Mathlib/Data/List/FinRange.lean @@ -5,6 +5,7 @@ Authors: Mario Carneiro, Kenny Lau, Kim Morrison, Alex Keizer -/ import Mathlib.Data.List.OfFn import Batteries.Data.List.Perm +import Batteries.Data.List.FinRange import Mathlib.Data.List.Nodup /-! @@ -25,9 +26,6 @@ variable {α : Type u} theorem finRange_eq_pmap_range (n : ℕ) : finRange n = (range n).pmap Fin.mk (by simp) := by apply List.ext_getElem <;> simp [finRange] -@[simp] -theorem finRange_zero : finRange 0 = [] := rfl - @[simp] theorem mem_finRange {n : ℕ} (a : Fin n) : a ∈ finRange n := by rw [finRange_eq_pmap_range] @@ -40,10 +38,6 @@ theorem nodup_finRange (n : ℕ) : (finRange n).Nodup := by rw [finRange_eq_pmap_range] exact (Pairwise.pmap (nodup_range n) _) fun _ _ _ _ => @Fin.ne_of_val_ne _ ⟨_, _⟩ ⟨_, _⟩ -@[simp] -theorem length_finRange (n : ℕ) : (finRange n).length = n := by - simp [finRange] - @[simp] theorem finRange_eq_nil {n : ℕ} : finRange n = [] ↔ n = 0 := by rw [← length_eq_zero, length_finRange] @@ -56,11 +50,6 @@ theorem pairwise_le_finRange (n : ℕ) : Pairwise (· ≤ ·) (finRange n) := by rw [finRange_eq_pmap_range] exact (List.pairwise_le_range n).pmap (by simp) (by simp) -@[simp] -theorem getElem_finRange {n : ℕ} {i : ℕ} (h) : - (finRange n)[i] = ⟨i, length_finRange n ▸ h⟩ := by - simp [finRange, getElem_range, getElem_pmap] - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/10756): new theorem theorem get_finRange {n : ℕ} {i : ℕ} (h) : (finRange n).get ⟨i, h⟩ = ⟨i, length_finRange n ▸ h⟩ := by @@ -88,11 +77,6 @@ theorem finRange_succ_eq_map (n : ℕ) : finRange n.succ = 0 :: (finRange n).map map_map] simp only [Function.comp_def, Fin.val_succ] -theorem finRange_succ (n : ℕ) : - finRange n.succ = (finRange n |>.map Fin.castSucc |>.concat (.last _)) := by - apply map_injective_iff.mpr Fin.val_injective - simp [range_succ, Function.comp_def] - -- Porting note: `map_nth_le` moved to `List.finRange_map_get` in Data.List.Range theorem ofFn_eq_pmap {n} {f : Fin n → α} : diff --git a/Mathlib/ModelTheory/Encoding.lean b/Mathlib/ModelTheory/Encoding.lean index 180cb86b870ed..e3832e469136b 100644 --- a/Mathlib/ModelTheory/Encoding.lean +++ b/Mathlib/ModelTheory/Encoding.lean @@ -82,8 +82,8 @@ 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_left, List.getElem_map, List.getElem_finRange] - simp only [length_map, length_finRange, i.2] + · simp only [length_map, length_finRange, is_lt, getElem_append_left, getElem_map, + getElem_finRange, cast_mk, Fin.eta] · simp only [length_map, length_finRange, drop_left'] /-- An encoding of terms as lists. -/ @@ -237,7 +237,7 @@ theorem listDecode_encode_list (l : List (Σn, L.BoundedFormula α n)) : 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_left, getElem_map] - · simp only [getElem_finRange, Fin.eta, Function.comp_apply, Sum.getLeft?] + · simp only [getElem_finRange, cast_mk, Fin.eta, Function.comp_apply, Sum.getLeft?_inl] · simp only [length_map, length_finRange, is_lt] rw [dif_pos] swap diff --git a/lake-manifest.json b/lake-manifest.json index 8152b35704a8a..a3160f181a6cb 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "7488499a8aad6ffada87ab6db73673d88dc04c97", + "rev": "f6d16c2a073e0385a362ad3e5d9e7e8260e298d6", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", diff --git a/lakefile.lean b/lakefile.lean index 300ff2124f428..cab011aba5693 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -23,8 +23,7 @@ require "leanprover-community" / "plausible" @ git "main" /-- These options are used * as `leanOptions`, prefixed by `` `weak``, so that `lake build` uses them; * as `moreServerArgs`, to set their default value in mathlib - (as well as `Archive`, `Counterexamples` and `test`). --/ + (as well as `Archive`, `Counterexamples` and `test`). -/ abbrev mathlibOnlyLinters : Array LeanOption := #[ ⟨`linter.docPrime, true⟩, ⟨`linter.hashCommand, true⟩, From 30c0e01daad4ff6def8c2d5dbc28b0cccd5d250e Mon Sep 17 00:00:00 2001 From: Daniel Carranza Date: Tue, 26 Nov 2024 22:40:37 +0000 Subject: [PATCH 328/829] feat(CategoryTheory/Enriched/Opposite): the opposite category of an enriched category (#19483) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit For a `V`-category `C`, define the opposite `V`-category as an instance on the type `Cᵒᵖ`, construct an equivalence between the underlying category `ForgetEnrichment V (Cᵒᵖ)` and the opposite category `(ForgetEnrichment V C)ᵒᵖ`, and construct an instance of `EnrichedOrdinaryCategory V (Cᵒᵖ)` in the case that `C` comes with an instance of `EnrichedOrdinaryCategory V C`. Co-authored-by: Daniel Carranza <61289743+daniel-carranza@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Enriched/Opposite.lean | 132 ++++++++++++++++++ 2 files changed, 133 insertions(+) create mode 100644 Mathlib/CategoryTheory/Enriched/Opposite.lean diff --git a/Mathlib.lean b/Mathlib.lean index 3741c55a437ae..c971dda66ba49 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1682,6 +1682,7 @@ import Mathlib.CategoryTheory.Endofunctor.Algebra import Mathlib.CategoryTheory.Endomorphism import Mathlib.CategoryTheory.Enriched.Basic import Mathlib.CategoryTheory.Enriched.FunctorCategory +import Mathlib.CategoryTheory.Enriched.Opposite import Mathlib.CategoryTheory.Enriched.Ordinary import Mathlib.CategoryTheory.EpiMono import Mathlib.CategoryTheory.EqToHom diff --git a/Mathlib/CategoryTheory/Enriched/Opposite.lean b/Mathlib/CategoryTheory/Enriched/Opposite.lean new file mode 100644 index 0000000000000..d25af18fc5ba9 --- /dev/null +++ b/Mathlib/CategoryTheory/Enriched/Opposite.lean @@ -0,0 +1,132 @@ +/- +Copyright (c) 2024 Daniel Carranza. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Daniel Carranza +-/ +import Mathlib.CategoryTheory.Enriched.Ordinary +import Mathlib.CategoryTheory.Monoidal.Braided.Basic + +/-! + +# The opposite category of an enriched category + +When a monoidal category `V` is braided, we may define the opposite `V`-category of a +`V`-category. The symmetry map is required to define the composition morphism. + +This file constructs the opposite `V`-category as an instance on the type `Cᵒᵖ` and constructs an +equivalence between + • `ForgetEnrichment V (Cᵒᵖ)`, the underlying category of the `V`-category `Cᵒᵖ`; and + • `(ForgetEnrichment V C)ᵒᵖ`, the opposite category of the underlying category of `C`. +We also show that if `C` is an enriched ordinary category (i.e. a category enriched in `V` +equipped with an identification `(X ⟶ Y) ≃ (𝟙_ V ⟶ (X ⟶[V] Y))`) then `Cᵒᵖ` is again +an enriched ordinary category. + +-/ + +universe v₁ u₁ v u + +namespace CategoryTheory + +open MonoidalCategory BraidedCategory + +variable (V : Type u₁) [Category.{v₁} V] [MonoidalCategory V] [BraidedCategory V] + +section + +variable (C : Type u) [EnrichedCategory V C] + +/-- For a `V`-category `C`, construct the opposite `V`-category structure on the type `Cᵒᵖ` +using the braiding in `V`. -/ +instance EnrichedCategory.opposite : EnrichedCategory V Cᵒᵖ where + Hom y x := EnrichedCategory.Hom x.unop y.unop + id x := EnrichedCategory.id x.unop + comp z y x := (β_ _ _).hom ≫ EnrichedCategory.comp (x.unop) (y.unop) (z.unop) + id_comp _ _ := by + simp only [braiding_naturality_left_assoc, braiding_tensorUnit_left, + Category.assoc, Iso.inv_hom_id_assoc] + exact EnrichedCategory.comp_id _ _ + comp_id _ _ := by + simp only [braiding_naturality_right_assoc, braiding_tensorUnit_right, + Category.assoc, Iso.inv_hom_id_assoc] + exact EnrichedCategory.id_comp _ _ + assoc _ _ _ _ := by + simp only [braiding_naturality_left_assoc, + MonoidalCategory.whiskerLeft_comp, Category.assoc] + rw [← EnrichedCategory.assoc] + simp only [braiding_tensor_left, Category.assoc, Iso.inv_hom_id_assoc, + braiding_naturality_right_assoc, braiding_tensor_right] + +end + +/-- Unfold the definition of composition in the enriched opposite category. -/ +@[reassoc] +lemma eComp_op_eq {C : Type u} [EnrichedCategory V C] (x y z : Cᵒᵖ) : + eComp V z y x = (β_ _ _).hom ≫ eComp V x.unop y.unop z.unop := + rfl + +/-- When composing a tensor product of morphisms with the `V`-composition morphism in `Cᵒᵖ`, +this re-writes the `V`-composition to be in `C` and moves the braiding to the left. -/ +@[reassoc] +lemma tensorHom_eComp_op_eq {C : Type u} [EnrichedCategory V C] {x y z : Cᵒᵖ} {v w : V} + (f : v ⟶ EnrichedCategory.Hom z y) (g : w ⟶ EnrichedCategory.Hom y x) : + (f ⊗ g) ≫ eComp V z y x = (β_ v w).hom ≫ (g ⊗ f) ≫ eComp V x.unop y.unop z.unop := by + rw [eComp_op_eq] + exact braiding_naturality_assoc f g _ + +-- This section establishes the equivalence on underlying categories +section + +open ForgetEnrichment + +variable (C : Type u) [EnrichedCategory V C] + +/-- The functor going from the underlying category of the enriched category `Cᵒᵖ` +to the opposite of the underlying category of the enriched category `C`. -/ +def forgetEnrichmentOppositeEquivalence.functor : + ForgetEnrichment V Cᵒᵖ ⥤ (ForgetEnrichment V C)ᵒᵖ where + obj x := x + map {x y} f := f.op + map_comp {x y z} f g := by + have : (f ≫ g) = homTo V (f ≫ g) := rfl + rw [this, forgetEnrichment_comp, Category.assoc, tensorHom_eComp_op_eq, + leftUnitor_inv_braiding_assoc, ← unitors_inv_equal, ← Category.assoc] + congr 1 + +/-- The functor going from the opposite of the underlying category of the enriched category `C` +to the underlying category of the enriched category `Cᵒᵖ`. -/ +def forgetEnrichmentOppositeEquivalence.inverse : + (ForgetEnrichment V C)ᵒᵖ ⥤ ForgetEnrichment V Cᵒᵖ where + obj x := x + map {x y} f := f.unop + map_comp {x y z} f g := by + have : g.unop ≫ f.unop = homTo V (g.unop ≫ f.unop) := rfl + dsimp + rw [this, forgetEnrichment_comp, Category.assoc, unitors_inv_equal, + ← leftUnitor_inv_braiding_assoc] + have : (β_ _ _).hom ≫ (homTo V g.unop ⊗ homTo V f.unop) ≫ + eComp V («to» V z.unop) («to» V y.unop) («to» V x.unop) = + ((homTo V f.unop) ⊗ (homTo V g.unop)) ≫ eComp V x y z := (tensorHom_eComp_op_eq V _ _).symm + rw [this, ← Category.assoc] + congr 1 + +/-- The equivalence between the underlying category of the enriched category `Cᵒᵖ` and +the opposite of the underlying category of the enriched category `C`. -/ +@[simps] +def forgetEnrichmentOppositeEquivalence : ForgetEnrichment V Cᵒᵖ ≌ (ForgetEnrichment V C)ᵒᵖ where + functor := forgetEnrichmentOppositeEquivalence.functor V C + inverse := forgetEnrichmentOppositeEquivalence.inverse V C + unitIso := NatIso.ofComponents (fun _ ↦ Iso.refl _) + counitIso := NatIso.ofComponents (fun _ ↦ Iso.refl _) + +/-- If `D` is an enriched ordinary category then `Dᵒᵖ` is an enriched ordinary category. -/ +instance EnrichedOrdinaryCategory.opposite {D : Type u} [Category.{v} D] + [EnrichedOrdinaryCategory V D] : EnrichedOrdinaryCategory V Dᵒᵖ where + homEquiv := Quiver.Hom.opEquiv.symm.trans homEquiv + homEquiv_id x := homEquiv_id (x.unop) + homEquiv_comp f g := by + simp only [unop_comp, tensorHom_eComp_op_eq, leftUnitor_inv_braiding_assoc, ← unitors_inv_equal] + exact homEquiv_comp g.unop f.unop + +end + +end CategoryTheory From 095a1c66b3101c42d97d114598a98d7e90916865 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Tue, 26 Nov 2024 22:40:38 +0000 Subject: [PATCH 329/829] feat: semiring with discrete PrimeSpectrum (#19496) If the prime spectrum of a commutative semiring R has discrete Zariski topology, then R is canonically isomorphic to the product of its localizations at the (finitely many) maximal ideals. This was previously only proven for commutative rings R; a different proof using the sheaf property is required to generalize it. --- .../GammaSpecAdjunction.lean | 2 +- .../PrimeSpectrum/Basic.lean | 84 +++++++------ Mathlib/RingTheory/Ideal/Basic.lean | 10 ++ .../Localization/Away/AdjoinRoot.lean | 2 +- .../RingTheory/Localization/Away/Basic.lean | 110 +++++++++++++++--- 5 files changed, 152 insertions(+), 56 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean index 87195c4877f61..188473407ba08 100644 --- a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean +++ b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean @@ -136,7 +136,7 @@ theorem toΓSpecCApp_iff -- Porting Note: Type class problem got stuck in `IsLocalization.Away.AwayMap.lift_comp` -- created instance manually. This replaces the `pick_goal` tactics have loc_inst := IsLocalization.to_basicOpen (Γ.obj (op X)) r - rw [← @IsLocalization.Away.AwayMap.lift_comp _ _ _ _ _ _ _ r loc_inst _ + rw [← @IsLocalization.Away.lift_comp _ _ _ _ _ _ _ r loc_inst _ (X.isUnit_res_toΓSpecMapBasicOpen r)] --pick_goal 5; exact is_localization.to_basic_open _ r constructor diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index 6a4899d2598b7..bb1699add44f8 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -9,7 +9,7 @@ import Mathlib.RingTheory.Ideal.Over import Mathlib.RingTheory.KrullDimension.Basic import Mathlib.RingTheory.LocalRing.ResidueField.Defs import Mathlib.RingTheory.LocalRing.RingHom.Basic -import Mathlib.RingTheory.Localization.Away.Lemmas +import Mathlib.RingTheory.Localization.Away.Basic import Mathlib.Tactic.StacksAttribute import Mathlib.Topology.KrullDimension import Mathlib.Topology.Sober @@ -530,18 +530,61 @@ theorem isLocalization_away_iff_atPrime_of_basicOpen_eq_singleton [Algebra R S] {f : R} {p : PrimeSpectrum R} (h : (basicOpen f).1 = {p}) : IsLocalization.Away f S ↔ IsLocalization.AtPrime S p.1 := have : IsLocalization.AtPrime (Localization.Away f) p.1 := by - refine .of_le_of_exists_dvd (Submonoid.powers f) _ + refine .of_le_of_exists_dvd (.powers f) _ (Submonoid.powers_le.mpr <| by apply h ▸ Set.mem_singleton p) fun r hr ↦ ?_ contrapose! hr simp_rw [← Ideal.mem_span_singleton] at hr have ⟨q, prime, le, disj⟩ := Ideal.exists_le_prime_disjoint (Ideal.span {r}) - (Submonoid.powers f) (Set.disjoint_right.mpr hr) + (.powers f) (Set.disjoint_right.mpr hr) have : ⟨q, prime⟩ ∈ (basicOpen f).1 := Set.disjoint_right.mp disj (Submonoid.mem_powers f) rw [h, Set.mem_singleton_iff] at this rw [← this] exact not_not.mpr (q.span_singleton_le_iff_mem.mp le) IsLocalization.isLocalization_iff_of_isLocalization _ _ (Localization.Away f) +variable [DiscreteTopology (PrimeSpectrum R)] + +variable (R) in +lemma _root_.RingHom.toLocalizationIsMaximal_surjective_of_discreteTopology : + Function.Surjective (RingHom.toLocalizationIsMaximal R) := fun x ↦ by + have (p : PrimeSpectrum R) : ∃ f, (basicOpen f : Set _) = {p} := + have ⟨_, ⟨f, rfl⟩, hpf, hfp⟩ := isTopologicalBasis_basic_opens.isOpen_iff.mp + (isOpen_discrete {p}) p rfl + ⟨f, hfp.antisymm <| Set.singleton_subset_iff.mpr hpf⟩ + choose f hf using this + let e := Equiv.ofInjective f fun p q eq ↦ Set.singleton_injective (hf p ▸ eq ▸ hf q) + have loc a : IsLocalization.AtPrime (Localization.Away a.1) (e.symm a).1 := + (isLocalization_away_iff_atPrime_of_basicOpen_eq_singleton <| hf _).mp <| by + simp_rw [Equiv.apply_ofInjective_symm]; infer_instance + let algE a := IsLocalization.algEquiv (e.symm a).1.primeCompl + (Localization.AtPrime (e.symm a).1) (Localization.Away a.1) + have span_eq : Ideal.span (Set.range f) = ⊤ := iSup_basicOpen_eq_top_iff.mp <| top_unique + fun p _ ↦ TopologicalSpace.Opens.mem_iSup.mpr ⟨p, (hf p).ge rfl⟩ + replace hf a : (basicOpen a.1 : Set _) = {e.symm a} := by + simp_rw [← hf, Equiv.apply_ofInjective_symm] + have := (discreteTopology_iff_finite_and_isPrime_imp_isMaximal.mp ‹_›).2 + obtain ⟨r, eq, -⟩ := Localization.existsUnique_algebraMap_eq_of_span_eq_top _ span_eq + (fun a ↦ algE a (x ⟨_, this _ inferInstance⟩)) fun a b ↦ by + obtain rfl | ne := eq_or_ne a b; · rfl + have ⟨n, hn⟩ : IsNilpotent (a * b : R) := (basicOpen_eq_bot_iff _).mp <| by + simp_rw [basicOpen_mul, SetLike.ext'_iff, TopologicalSpace.Opens.coe_inf, hf] + exact bot_unique (fun _ ⟨ha, hb⟩ ↦ ne <| e.symm.injective (ha.symm.trans hb)) + have := IsLocalization.subsingleton (M := .powers (a * b : R)) + (S := Localization.Away (a * b : R)) <| hn ▸ ⟨n, rfl⟩ + apply Subsingleton.elim + refine ⟨r, funext fun I ↦ ?_⟩ + have := eq (e ⟨I, I.2.isPrime⟩) + rwa [← AlgEquiv.symm_apply_eq, AlgEquiv.commutes, e.symm_apply_apply] at this + +/-- If the prime spectrum of a commutative semiring R has discrete Zariski topology, then R is +canonically isomorphic to the product of its localizations at the (finitely many) maximal ideals. -/ +@[stacks 00JA +"See also `PrimeSpectrum.discreteTopology_iff_finite_isMaximal_and_sInf_le_nilradical`."] +def _root_.RingHom.toLocalizationIsMaximalEquiv : R ≃+* + Π I : {I : Ideal R // I.IsMaximal}, haveI : I.1.IsMaximal := I.2; Localization.AtPrime I.1 := + .ofBijective _ ⟨RingHom.toLocalizationIsMaximal_injective R, + RingHom.toLocalizationIsMaximal_surjective_of_discreteTopology R⟩ + end BasicOpen section Order @@ -1102,39 +1145,4 @@ lemma isIdempotentElemEquivClopens_symm_sup (s₁ s₂ : Clopens (PrimeSpectrum end PrimeSpectrum -variable [DiscreteTopology (PrimeSpectrum R)] -open PrimeSpectrum - -variable (R) in -lemma RingHom.toLocalizationIsMaximal_surjective_of_discreteTopology : - Function.Surjective (RingHom.toLocalizationIsMaximal R) := fun x ↦ by - let idem I := isIdempotentElemEquivClopens (R := R).symm ⟨{I}, isClopen_discrete _⟩ - let ideal I := Ideal.span {1 - (idem I).1} - let toSpec (I : {I : Ideal R | I.IsMaximal}) : PrimeSpectrum R := ⟨I.1, I.2.isPrime⟩ - have loc I : IsLocalization.AtPrime (R ⧸ ideal I) I.1 := by - rw [← isLocalization_away_iff_atPrime_of_basicOpen_eq_singleton] - exacts [IsLocalization.Away.quotient_of_isIdempotentElem (idem I).2, - congr_arg (·.1) (basicOpen_isIdempotentElemEquivClopens_symm ⟨{I}, isClopen_discrete _⟩)] - let equiv I := IsLocalization.algEquiv I.1.primeCompl (Localization.AtPrime I.1) (R ⧸ ideal I) - have := (discreteTopology_iff_finite_isMaximal_and_sInf_le_nilradical.mp ‹_›).1 - have ⟨r, hr⟩ := Ideal.pi_quotient_surjective ?_ fun I ↦ equiv (toSpec I) (x I) - · refine ⟨r, funext fun I ↦ (equiv <| toSpec I).injective ?_⟩ - rw [← hr]; exact (equiv _).commutes r - refine fun I J ne ↦ Ideal.isCoprime_iff_exists.mpr ?_ - have := ((idem <| toSpec I).2.mul (idem <| toSpec J).2).eq_zero_of_isNilpotent <| by - simp_rw [← basicOpen_eq_bot_iff, basicOpen_mul, SetLike.ext'_iff, idem, - TopologicalSpace.Opens.coe_inf, basicOpen_isIdempotentElemEquivClopens_symm] - exact Set.singleton_inter_eq_empty.mpr fun h ↦ ne (Subtype.ext <| congr_arg (·.1) h) - simp_rw [ideal, Ideal.mem_span_singleton', exists_exists_eq_and] - exact ⟨1, idem (toSpec I), by simpa [mul_sub]⟩ - -/-- If the prime spectrum of a commutative ring R has discrete Zariski topology, then R is -canonically isomorphic to the product of its localizations at the (finitely many) maximal ideals. -/ -@[stacks 00JA -"See also `PrimeSpectrum.discreteTopology_iff_finite_isMaximal_and_sInf_le_nilradical`."] -def RingHom.toLocalizationIsMaximalEquiv : R ≃+* - Π I : {I : Ideal R // I.IsMaximal}, haveI : I.1.IsMaximal := I.2; Localization.AtPrime I.1 := - .ofBijective _ ⟨RingHom.toLocalizationIsMaximal_injective R, - RingHom.toLocalizationIsMaximal_surjective_of_discreteTopology R⟩ - end Idempotent diff --git a/Mathlib/RingTheory/Ideal/Basic.lean b/Mathlib/RingTheory/Ideal/Basic.lean index 0c7c7a913ae1b..9bcac7b686908 100644 --- a/Mathlib/RingTheory/Ideal/Basic.lean +++ b/Mathlib/RingTheory/Ideal/Basic.lean @@ -150,6 +150,16 @@ theorem span_pow_eq_top (s : Set α) (hs : span s = ⊤) (n : ℕ) : rw [mul_pow, mem_span_singleton] exact ⟨f x ^ (n + 1), mul_comm _ _⟩ +theorem span_range_pow_eq_top (s : Set α) (hs : span s = ⊤) (n : s → ℕ) : + span (Set.range fun x ↦ x.1 ^ n x) = ⊤ := by + have ⟨t, hts, mem⟩ := Submodule.mem_span_finite_of_mem_span ((eq_top_iff_one _).mp hs) + refine top_unique ((span_pow_eq_top _ ((eq_top_iff_one _).mpr mem) <| + t.attach.sup fun x ↦ n ⟨x, hts x.2⟩).ge.trans <| span_le.mpr ?_) + rintro _ ⟨x, hxt, rfl⟩ + rw [← Nat.sub_add_cancel (Finset.le_sup <| t.mem_attach ⟨x, hxt⟩)] + simp_rw [pow_add] + exact mul_mem_left _ _ (subset_span ⟨_, rfl⟩) + end Ideal end CommSemiring diff --git a/Mathlib/RingTheory/Localization/Away/AdjoinRoot.lean b/Mathlib/RingTheory/Localization/Away/AdjoinRoot.lean index 5211fab000937..4061c6c8231de 100644 --- a/Mathlib/RingTheory/Localization/Away/AdjoinRoot.lean +++ b/Mathlib/RingTheory/Localization/Away/AdjoinRoot.lean @@ -28,7 +28,7 @@ noncomputable def Localization.awayEquivAdjoin (r : R) : Away r ≃ₐ[R] Adjoin (isUnit_of_mul_eq_one ((algebraMap R (AdjoinRoot (C r * X - 1))) r) (root (C r * X - 1)) (root_isInv r)) with commutes' := - IsLocalization.Away.AwayMap.lift_eq r (isUnit_of_mul_eq_one _ _ <| root_isInv r) } + IsLocalization.Away.lift_eq r (isUnit_of_mul_eq_one _ _ <| root_isInv r) } (liftHom _ (IsLocalization.Away.invSelf r) <| by simp only [map_sub, map_mul, aeval_C, aeval_X, IsLocalization.Away.mul_invSelf, aeval_one, sub_self]) diff --git a/Mathlib/RingTheory/Localization/Away/Basic.lean b/Mathlib/RingTheory/Localization/Away/Basic.lean index ae069a4b51072..259a7cba398d2 100644 --- a/Mathlib/RingTheory/Localization/Away/Basic.lean +++ b/Mathlib/RingTheory/Localization/Away/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Mario Carneiro, Johan Commelin, Amelia Livingston, Anne Baanen -/ import Mathlib.GroupTheory.MonoidLocalization.Away +import Mathlib.Algebra.Algebra.Pi import Mathlib.RingTheory.Ideal.Maps import Mathlib.RingTheory.Localization.Basic import Mathlib.RingTheory.UniqueFactorizationDomain.Multiplicity @@ -123,6 +124,9 @@ 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⟩ +lemma isUnit_of_dvd {r : R} (h : r ∣ x) : IsUnit (algebraMap R S r) := + isUnit_of_dvd_unit (map_dvd _ h) (algebraMap_isUnit x) + variable {g : R →+* P} /-- Given `x : R`, a localization map `F : R →+* S` away from `x`, and a map of `CommSemiring`s @@ -136,20 +140,33 @@ noncomputable def lift (hg : IsUnit (g x)) : S →+* P := exact IsUnit.map (powMonoidHom n : P →* P) hg @[simp] -theorem AwayMap.lift_eq (hg : IsUnit (g x)) (a : R) : lift x hg ((algebraMap R S) a) = g a := +theorem lift_eq (hg : IsUnit (g x)) (a : R) : lift x hg (algebraMap R S a) = g a := IsLocalization.lift_eq _ _ @[simp] -theorem AwayMap.lift_comp (hg : IsUnit (g x)) : (lift x hg).comp (algebraMap R S) = g := +theorem lift_comp (hg : IsUnit (g x)) : (lift x hg).comp (algebraMap R S) = g := IsLocalization.lift_comp _ +@[deprecated (since := "2024-11-25")] alias AwayMap.lift_eq := lift_eq +@[deprecated (since := "2024-11-25")] alias AwayMap.lift_comp := lift_comp + +/-- Given `x y : R` and localizations `S`, `P` away from `x` and `y * x` +respectively, the homomorphism induced from `S` to `P`. -/ +noncomputable def awayToAwayLeft (y : R) [Algebra R P] [IsLocalization.Away (y * x) P] : S →+* P := + lift x <| isUnit_of_dvd (y * x) (dvd_mul_left _ _) + /-- Given `x y : R` and localizations `S`, `P` away from `x` and `x * y` respectively, the homomorphism induced from `S` to `P`. -/ noncomputable def awayToAwayRight (y : R) [Algebra R P] [IsLocalization.Away (x * y) P] : S →+* P := - lift x <| - show IsUnit ((algebraMap R P) x) from - isUnit_of_mul_eq_one ((algebraMap R P) x) (mk' P y ⟨x * y, Submonoid.mem_powers _⟩) <| by - rw [mul_mk'_eq_mk'_of_mul, mk'_self] + lift x <| isUnit_of_dvd (x * y) (dvd_mul_right _ _) + +theorem awayToAwayLeft_eq (y : R) [Algebra R P] [IsLocalization.Away (y * x) P] (a : R) : + awayToAwayLeft x y (algebraMap R S a) = algebraMap R P a := + lift_eq _ _ _ + +theorem awayToAwayRight_eq (y : R) [Algebra R P] [IsLocalization.Away (x * y) P] (a : R) : + awayToAwayRight x y (algebraMap R S a) = algebraMap R P a := + lift_eq _ _ _ variable (S) (Q : Type*) [CommSemiring Q] [Algebra P Q] @@ -164,10 +181,10 @@ noncomputable def map (f : R →+* P) (r : R) [IsLocalization.Away r S] section Algebra -variable {A : Type*} [CommRing A] [Algebra R A] -variable {B : Type*} [CommRing B] [Algebra R B] -variable (Aₚ : Type*) [CommRing Aₚ] [Algebra A Aₚ] [Algebra R Aₚ] [IsScalarTower R A Aₚ] -variable (Bₚ : Type*) [CommRing Bₚ] [Algebra B Bₚ] [Algebra R Bₚ] [IsScalarTower R B Bₚ] +variable {A : Type*} [CommSemiring A] [Algebra R A] +variable {B : Type*} [CommSemiring B] [Algebra R B] +variable (Aₚ : Type*) [CommSemiring Aₚ] [Algebra A Aₚ] [Algebra R Aₚ] [IsScalarTower R A Aₚ] +variable (Bₚ : Type*) [CommSemiring Bₚ] [Algebra B Bₚ] [Algebra R Bₚ] [IsScalarTower R B Bₚ] instance {f : A →+* B} (a : A) [Away (f a) Bₚ] : IsLocalization (.map f (.powers a)) Bₚ := by simpa @@ -279,7 +296,7 @@ noncomputable def atUnit (x : R) (e : IsUnit x) [IsLocalization.Away x S] : R noncomputable def atOne [IsLocalization.Away (1 : R) S] : R ≃ₐ[R] S := @atUnit R _ S _ _ (1 : R) isUnit_one _ -theorem away_of_isUnit_of_bijective {R : Type*} (S : Type*) [CommRing R] [CommRing S] +theorem away_of_isUnit_of_bijective {R : Type*} (S : Type*) [CommSemiring R] [CommSemiring S] [Algebra R S] {r : R} (hr : IsUnit r) (H : Function.Bijective (algebraMap R S)) : IsLocalization.Away r S := { map_units' := by @@ -364,28 +381,88 @@ noncomputable abbrev awayMap (f : R →+* P) (r : R) : Localization.Away r →+* Localization.Away (f r) := IsLocalization.Away.map _ _ f r -variable {A : Type*} [CommRing A] [Algebra R A] -variable {B : Type*} [CommRing B] [Algebra R B] +variable {A : Type*} [CommSemiring A] [Algebra R A] +variable {B : Type*} [CommSemiring B] [Algebra R B] /-- Given a map `f : A →ₐ[R] B` and an element `a : A`, we may construct a map `Aₐ →ₐ[R] Bₐ`. -/ noncomputable abbrev awayMapₐ (f : A →ₐ[R] B) (a : A) : Localization.Away a →ₐ[R] Localization.Away (f a) := IsLocalization.Away.mapₐ _ _ f a +theorem algebraMap_injective_of_span_eq_top (s : Set R) (span_eq : Ideal.span s = ⊤) : + Function.Injective (algebraMap R <| Π a : s, Away a.1) := fun x y eq ↦ by + suffices Module.eqIdeal R x y = ⊤ by simpa [Module.eqIdeal] using (Ideal.eq_top_iff_one _).mp this + by_contra ne + have ⟨r, hrs, disj⟩ := Ideal.exists_disjoint_powers_of_span_eq_top s span_eq _ ne + let r : s := ⟨r, hrs⟩ + have ⟨⟨_, n, rfl⟩, eq⟩ := (IsLocalization.eq_iff_exists (.powers r.1) _).mp (congr_fun eq r) + exact Set.disjoint_left.mp disj eq ⟨n, rfl⟩ + +/-- The sheaf condition for the structure sheaf on `Spec R` +for a covering of the whole prime spectrum by basic opens. -/ +theorem existsUnique_algebraMap_eq_of_span_eq_top (s : Set R) (span_eq : Ideal.span s = ⊤) + (f : Π a : s, Away a.1) (h : ∀ a b : s, + Away.awayToAwayRight (P := Away (a * b : R)) a.1 b (f a) = Away.awayToAwayLeft b.1 a (f b)) : + ∃! r : R, ∀ a : s, algebraMap R _ r = f a := by + have mem := (Ideal.eq_top_iff_one _).mp span_eq; clear span_eq + wlog finset_eq : ∃ t : Finset R, t = s generalizing s + · have ⟨t, hts, mem⟩ := Submodule.mem_span_finite_of_mem_span mem + have ⟨r, eq, uniq⟩ := this t (fun a ↦ f ⟨a, hts a.2⟩) + (fun a b ↦ h ⟨a, hts a.2⟩ ⟨b, hts b.2⟩) mem ⟨_, rfl⟩ + refine ⟨r, fun a ↦ ?_, fun _ eq ↦ uniq _ fun a ↦ eq ⟨a, hts a.2⟩⟩ + replace hts := Set.insert_subset a.2 hts + classical + have ⟨r', eq, _⟩ := this ({a.1} ∪ t) (fun a ↦ f ⟨a, hts a.2⟩) (fun a b ↦ + h ⟨a, hts a.2⟩ ⟨b, hts b.2⟩) (Ideal.span_mono (fun _ ↦ .inr) mem) ⟨{a.1} ∪ t, by simp⟩ + exact (congr_arg _ (uniq _ fun b ↦ eq ⟨b, .inr b.2⟩).symm).trans (eq ⟨a, .inl rfl⟩) + have span_eq := (Ideal.eq_top_iff_one _).mpr mem + refine exists_unique_of_exists_of_unique ?_ fun x y hx hy ↦ + algebraMap_injective_of_span_eq_top s span_eq (funext fun a ↦ (hx a).trans (hy a).symm) + obtain ⟨s, rfl⟩ := finset_eq + choose n r eq using fun a ↦ Away.surj a.1 (f a) + let N := s.attach.sup n + let r a := a ^ (N - n a) * r a + have eq a : f a * algebraMap R _ (a ^ N) = algebraMap R _ (r a) := by + rw [map_mul, ← eq, mul_left_comm, ← map_pow, ← map_mul, ← pow_add, + Nat.sub_add_cancel (Finset.le_sup <| s.mem_attach a)] + have eq2 a b : ∃ N', (a * b) ^ N' * (r a * b ^ N) = (a * b) ^ N' * (r b * a ^ N) := + Away.exists_of_eq (S := Away (a * b : R)) _ <| by + simp_rw [map_mul, ← Away.awayToAwayRight_eq (S := Away a.1) a.1 b (r a), + ← Away.awayToAwayLeft_eq (S := Away b.1) b.1 a (r b), ← eq, map_mul, + Away.awayToAwayRight_eq, Away.awayToAwayLeft_eq, h, mul_assoc, ← map_mul, mul_comm] + choose N' hN' using eq2 + let N' := (s ×ˢ s).attach.sup fun a ↦ N' + ⟨_, (Finset.mem_product.mp a.2).1⟩ ⟨_, (Finset.mem_product.mp a.2).2⟩ + have eq2 a b : (a * b) ^ N' * (r a * b ^ N) = (a * b) ^ N' * (r b * a ^ N) := by + dsimp only [N']; rw [← Nat.sub_add_cancel (Finset.le_sup <| (Finset.mem_attach _ ⟨⟨a, b⟩, + Finset.mk_mem_product a.2 b.2⟩)), pow_add, mul_assoc, hN', ← mul_assoc] + let N := N' + N + let r a := a ^ N' * r a + have eq a : f a * algebraMap R _ (a ^ N) = algebraMap R _ (r a) := by + rw [map_mul, ← eq, mul_left_comm, ← map_mul, ← pow_add] + have eq2 a b : r a * b ^ N = r b * a ^ N := by + rw [pow_add, mul_mul_mul_comm, ← mul_pow, eq2, + mul_comm a.1, mul_pow, mul_mul_mul_comm, ← pow_add] + have ⟨c, eq1⟩ := (mem_span_range_iff_exists_fun _).mp <| + (Ideal.eq_top_iff_one _).mp <| (Set.image_eq_range _ _ ▸ Ideal.span_pow_eq_top _ span_eq N) + refine ⟨∑ b, c b * r b, fun a ↦ ((Away.algebraMap_isUnit a.1).pow N).mul_left_inj.mp ?_⟩ + simp_rw [← map_pow, eq, ← map_mul, Finset.sum_mul, mul_assoc, eq2 _ a, mul_left_comm (c _), + ← Finset.mul_sum, ← smul_eq_mul (a := c _), eq1, mul_one] + end Localization end CommSemiring open Localization -variable {R : Type*} [CommRing R] +variable {R : Type*} [CommSemiring R] section NumDen open IsLocalization variable (x : R) -variable (B : Type*) [CommRing B] [Algebra R B] [IsLocalization.Away x B] +variable (B : Type*) [CommSemiring B] [Algebra R B] [IsLocalization.Away x B] /-- `selfZPow x (m : ℤ)` is `x ^ m` as an element of the localization away from `x`. -/ noncomputable def selfZPow (m : ℤ) : B := @@ -478,7 +555,8 @@ theorem selfZPow_pow_sub (a : R) (b : B) (m d : ℤ) : simp only at this rwa [mul_comm _ b, mul_assoc b _ _, selfZPow_mul_neg, mul_one] at this -variable [IsDomain R] [WfDvdMonoid R] +variable {R : Type*} [CommRing R] (x : R) (B : Type*) [CommRing B] +variable [Algebra R B] [IsLocalization.Away x B] [IsDomain R] [WfDvdMonoid R] theorem exists_reduced_fraction' {b : B} (hb : b ≠ 0) (hx : Irreducible x) : ∃ (a : R) (n : ℤ), ¬x ∣ a ∧ selfZPow x B n * algebraMap R B a = b := by From de39dd8ecfea9b240dede98cdbbedf44f906a37f Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 26 Nov 2024 22:40:39 +0000 Subject: [PATCH 330/829] feat: zulip emoji reactions on public and `mathlib reviewers` channels (#19522) Updates the action that places emoji reactions triggered by `bors x`: now the reactions should be added to all public channels and to the `mathlib reviewers` channel. --- scripts/zulip_emoji_merge_delegate.py | 29 ++++++++++++++++++++++----- 1 file changed, 24 insertions(+), 5 deletions(-) diff --git a/scripts/zulip_emoji_merge_delegate.py b/scripts/zulip_emoji_merge_delegate.py index e572aeb90e8d3..8fa4e33e74f4c 100644 --- a/scripts/zulip_emoji_merge_delegate.py +++ b/scripts/zulip_emoji_merge_delegate.py @@ -23,15 +23,34 @@ site=ZULIP_SITE ) -# Fetch the last 200 messages -response = client.get_messages({ +# Fetch the messages containing the PR number from the public channels. +# There does not seem to be a way to search simultaneously public and private channels. +public_response = client.get_messages({ "anchor": "newest", - "num_before": 200, + "num_before": 5000, "num_after": 0, - "narrow": [{"operator": "channel", "operand": "PR reviews"}], + "narrow": [ + {"operator": "channels", "operand": "public"}, + {"operator": "search", "operand": f'#{PR_NUMBER}'}, + ], }) -messages = response['messages'] +# Fetch the messages containing the PR number from the `mathlib reviewers` channel +# There does not seem to be a way to search simultaneously public and private channels. +reviewers_response = client.get_messages({ + "anchor": "newest", + "num_before": 5000, + "num_after": 0, + "narrow": [ + {"operator": "channel", "operand": "mathlib reviewers"}, + {"operator": "search", "operand": f'#{PR_NUMBER}'}, + ], +}) + +print(f"public_response:{public_response}") +print(f"reviewers_response:{reviewers_response}") + +messages = (public_response['messages']) + (reviewers_response['messages']) pr_pattern = re.compile(f'https://github.com/leanprover-community/mathlib4/pull/{PR_NUMBER}') From 84b098419d680a531472f202a5ce5234bec118b6 Mon Sep 17 00:00:00 2001 From: peabrainiac Date: Tue, 26 Nov 2024 23:03:57 +0000 Subject: [PATCH 331/829] feat(Topology/Compactness): delta-generated spaces (#19431) Introduces delta-generated topological spaces and shows that they are sequential, locally path-connected and closed under disjoint unions and quotients. Co-authored-by: peabrainiac <43812953+peabrainiac@users.noreply.github.com> --- Mathlib.lean | 1 + .../Compactness/DeltaGeneratedSpace.lean | 183 ++++++++++++++++++ 2 files changed, 184 insertions(+) create mode 100644 Mathlib/Topology/Compactness/DeltaGeneratedSpace.lean diff --git a/Mathlib.lean b/Mathlib.lean index c971dda66ba49..1fe86e9134e3b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -5033,6 +5033,7 @@ import Mathlib.Topology.Compactification.OnePoint import Mathlib.Topology.Compactification.OnePointEquiv import Mathlib.Topology.Compactness.Compact import Mathlib.Topology.Compactness.CompactlyGeneratedSpace +import Mathlib.Topology.Compactness.DeltaGeneratedSpace import Mathlib.Topology.Compactness.Exterior import Mathlib.Topology.Compactness.Lindelof import Mathlib.Topology.Compactness.LocallyCompact diff --git a/Mathlib/Topology/Compactness/DeltaGeneratedSpace.lean b/Mathlib/Topology/Compactness/DeltaGeneratedSpace.lean new file mode 100644 index 0000000000000..c7c9e34c6a6ae --- /dev/null +++ b/Mathlib/Topology/Compactness/DeltaGeneratedSpace.lean @@ -0,0 +1,183 @@ +/- +Copyright (c) 2024 Ben Eltschig. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Ben Eltschig +-/ +import Mathlib.Analysis.Convex.Normed + +/-! +# Delta-generated topological spaces + +This file defines delta-generated spaces, as topological spaces whose topology is coinduced by all +maps from euclidean spaces into them. This is the strongest topological property that holds for +all CW-complexes and is closed under quotients and disjoint unions; every delta-generated space is +locally path-connected, sequential and in particular compactly generated. + +See https://ncatlab.org/nlab/show/Delta-generated+topological+space. + +Adapted from `Mathlib.Topology.Compactness.CompactlyGeneratedSpace`. + +## TODO +* All locally path-connected first-countable spaces are delta-generated - in particular, all normed + spaces and convex subsets thereof, like the standard simplices and the unit interval. +* Delta-generated spaces are equivalently generated by the simplices Δⁿ. +* Delta-generated spaces are equivalently generated by the unit interval I. +-/ + +variable {X Y : Type*} [tX : TopologicalSpace X] [tY : TopologicalSpace Y] + +open TopologicalSpace Topology + +/-- The topology coinduced by all maps from ℝⁿ into a space. -/ +def TopologicalSpace.deltaGenerated (X : Type*) [TopologicalSpace X] : TopologicalSpace X := + ⨆ f : (n : ℕ) × C(((Fin n) → ℝ), X), coinduced f.2 inferInstance + +/-- The delta-generated topology is also coinduced by a single map out of a sigma type. -/ +lemma deltaGenerated_eq_coinduced : deltaGenerated X = coinduced + (fun x : (f : (n : ℕ) × C(Fin n → ℝ, X)) × (Fin f.1 → ℝ) ↦ x.1.2 x.2) inferInstance := by + rw [deltaGenerated, instTopologicalSpaceSigma, coinduced_iSup]; rfl + +/-- The delta-generated topology is at least as fine as the original one. -/ +lemma deltaGenerated_le : deltaGenerated X ≤ tX := + iSup_le_iff.mpr fun f ↦ f.2.continuous.coinduced_le + +/-- A set is open in `deltaGenerated X` iff all its preimages under continuous functions ℝⁿ → X are + open. -/ +lemma isOpen_deltaGenerated_iff {u : Set X} : + IsOpen[deltaGenerated X] u ↔ ∀ n (p : C(Fin n → ℝ, X)), IsOpen (p ⁻¹' u) := by + simp_rw [deltaGenerated, isOpen_iSup_iff, isOpen_coinduced, Sigma.forall] + +/-- A map from ℝⁿ to X is continuous iff it is continuous regarding the + delta-generated topology on X. Outside of this file, use the more general + `continuous_to_deltaGenerated` instead. -/ +private lemma continuous_euclidean_to_deltaGenerated {n : ℕ} {f : (Fin n → ℝ) → X} : + Continuous[_, deltaGenerated X] f ↔ Continuous f := by + simp_rw [continuous_iff_coinduced_le] + refine ⟨fun h ↦ h.trans deltaGenerated_le, fun h ↦ ?_⟩ + simp_rw [deltaGenerated] + exact le_iSup_of_le (i := ⟨n, f, continuous_iff_coinduced_le.mpr h⟩) le_rfl + +/-- `deltaGenerated` is idempotent as a function `TopologicalSpace X → TopologicalSpace X`. -/ +lemma deltaGenerated_deltaGenerated_eq : + @deltaGenerated X (deltaGenerated X) = deltaGenerated X := by + ext u; simp_rw [isOpen_deltaGenerated_iff]; refine forall_congr' fun n ↦ ?_ + -- somewhat awkward because `ContinuousMap` doesn't play well with multiple topologies. + refine ⟨fun h p ↦ h <| @ContinuousMap.mk _ _ _ (_) p ?_, fun h p ↦ h ⟨p, ?_⟩⟩ + · exact continuous_euclidean_to_deltaGenerated.mpr p.2 + · exact continuous_euclidean_to_deltaGenerated.mp <| @ContinuousMap.continuous_toFun _ _ _ (_) p + +/-- A space is delta-generated if its topology is equal to the delta-generated topology, i.e. + coinduced by all continuous maps ℝⁿ → X. Since the delta-generated topology is always finer + than the original one, it suffices to show that it is also coarser. -/ +class DeltaGeneratedSpace (X : Type*) [t : TopologicalSpace X] : Prop where + le_deltaGenerated : t ≤ deltaGenerated X + +lemma eq_deltaGenerated [DeltaGeneratedSpace X] : tX = deltaGenerated X := + eq_of_le_of_le DeltaGeneratedSpace.le_deltaGenerated deltaGenerated_le + +/-- A subset of a delta-generated space is open iff its preimage is open for every + continuous map from ℝⁿ to X. -/ +lemma DeltaGeneratedSpace.isOpen_iff [DeltaGeneratedSpace X] {u : Set X} : + IsOpen u ↔ ∀ (n : ℕ) (p : ContinuousMap ((Fin n) → ℝ) X), IsOpen (p ⁻¹' u) := by + nth_rewrite 1 [eq_deltaGenerated (X := X)]; exact isOpen_deltaGenerated_iff + +/-- A map out of a delta-generated space is continuous iff it preserves continuity of maps + from ℝⁿ into X. -/ +lemma DeltaGeneratedSpace.continuous_iff [DeltaGeneratedSpace X] {f : X → Y} : + Continuous f ↔ ∀ (n : ℕ) (p : C(((Fin n) → ℝ), X)), Continuous (f ∘ p) := by + simp_rw [continuous_iff_coinduced_le] + nth_rewrite 1 [eq_deltaGenerated (X := X), deltaGenerated] + simp [coinduced_compose] + constructor + · intro h n p; apply h ⟨n, p⟩ + · rintro h ⟨n, p⟩; apply h n p + +/-- A map out of a delta-generated space is continuous iff it is continuous with respect + to the delta-generated topology on the codomain. -/ +lemma continuous_to_deltaGenerated [DeltaGeneratedSpace X] {f : X → Y} : + Continuous[_, deltaGenerated Y] f ↔ Continuous f := by + simp_rw [DeltaGeneratedSpace.continuous_iff, continuous_euclidean_to_deltaGenerated] + +/-- The delta-generated topology on `X` does in fact turn `X` into a delta-generated space. -/ +lemma deltaGeneratedSpace_deltaGenerated {X : Type*} {t : TopologicalSpace X} : + @DeltaGeneratedSpace X (@deltaGenerated X t) := by + let _ := @deltaGenerated X t; constructor; rw [@deltaGenerated_deltaGenerated_eq X t] + +lemma deltaGenerated_mono {X : Type*} {t₁ t₂ : TopologicalSpace X} (h : t₁ ≤ t₂) : + @deltaGenerated X t₁ ≤ @deltaGenerated X t₂ := by + rw [← continuous_id_iff_le, @continuous_to_deltaGenerated _ _ + (@deltaGenerated X t₁) t₂ deltaGeneratedSpace_deltaGenerated id] + exact continuous_id_iff_le.2 <| (@deltaGenerated_le X t₁).trans h + +namespace DeltaGeneratedSpace + +/-- Type synonym to be equipped with the delta-generated topology. -/ +def of (X : Type*) := X + +instance : TopologicalSpace (of X) := deltaGenerated X + +instance : DeltaGeneratedSpace (of X) := + deltaGeneratedSpace_deltaGenerated + +/-- The natural map from `DeltaGeneratedSpace.of X` to `X`. -/ +def counit : (of X) → X := id + +lemma continuous_counit : Continuous (counit : _ → X) := by + rw [continuous_iff_coinduced_le]; exact deltaGenerated_le + +/-- Delta-generated spaces are locally path-connected. -/ +instance [DeltaGeneratedSpace X] : LocPathConnectedSpace X := by + rw [eq_deltaGenerated (X := X), deltaGenerated_eq_coinduced] + exact LocPathConnectedSpace.coinduced _ + +/-- Delta-generated spaces are sequential. -/ +instance [DeltaGeneratedSpace X] : SequentialSpace X := by + rw [eq_deltaGenerated (X := X)] + exact SequentialSpace.iSup fun p ↦ SequentialSpace.coinduced p.2 + +end DeltaGeneratedSpace + +omit tY in +/-- Any topology coinduced by a delta-generated topology is delta-generated. -/ +lemma DeltaGeneratedSpace.coinduced [DeltaGeneratedSpace X] (f : X → Y) : + @DeltaGeneratedSpace Y (tX.coinduced f) := + let _ := tX.coinduced f + ⟨(continuous_to_deltaGenerated.2 continuous_coinduced_rng).coinduced_le⟩ + +/-- Suprema of delta-generated topologies are delta-generated. -/ +protected lemma DeltaGeneratedSpace.iSup {X : Type*} {ι : Sort*} {t : ι → TopologicalSpace X} + (h : ∀ i, @DeltaGeneratedSpace X (t i)) : @DeltaGeneratedSpace X (⨆ i, t i) := + let _ := ⨆ i, t i + ⟨iSup_le_iff.2 fun i ↦ (h i).le_deltaGenerated.trans <| deltaGenerated_mono <| le_iSup t i⟩ + +/-- Suprema of delta-generated topologies are delta-generated. -/ +protected lemma DeltaGeneratedSpace.sup {X : Type*} {t₁ t₂ : TopologicalSpace X} + (h₁ : @DeltaGeneratedSpace X t₁) (h₂ : @DeltaGeneratedSpace X t₂) : + @DeltaGeneratedSpace X (t₁ ⊔ t₂) := by + rw [sup_eq_iSup] + exact .iSup <| Bool.forall_bool.2 ⟨h₂, h₁⟩ + +/-- Quotients of delta-generated spaces are delta-generated. -/ +lemma Topology.IsQuotientMap.deltaGeneratedSpace [DeltaGeneratedSpace X] + {f : X → Y} (h : IsQuotientMap f) : DeltaGeneratedSpace Y := + h.2 ▸ DeltaGeneratedSpace.coinduced f + +/-- Quotients of delta-generated spaces are delta-generated. -/ +instance Quot.deltaGeneratedSpace [DeltaGeneratedSpace X] {r : X → X → Prop} : + DeltaGeneratedSpace (Quot r) := + isQuotientMap_quot_mk.deltaGeneratedSpace + +/-- Quotients of delta-generated spaces are delta-generated. -/ +instance Quotient.deltaGeneratedSpace [DeltaGeneratedSpace X] {s : Setoid X} : + DeltaGeneratedSpace (Quotient s) := + isQuotientMap_quotient_mk'.deltaGeneratedSpace + +/-- Disjoint unions of delta-generated spaces are delta-generated. -/ +instance Sum.deltaGeneratedSpace [DeltaGeneratedSpace X] [DeltaGeneratedSpace Y] : + DeltaGeneratedSpace (X ⊕ Y) := + DeltaGeneratedSpace.sup (.coinduced Sum.inl) (.coinduced Sum.inr) + +/-- Disjoint unions of delta-generated spaces are delta-generated. -/ +instance Sigma.deltaGeneratedSpace {ι : Type*} {X : ι → Type*} [∀ i, TopologicalSpace (X i)] + [∀ i, DeltaGeneratedSpace (X i)] : DeltaGeneratedSpace (Σ i, X i) := + .iSup fun _ ↦ .coinduced _ From 1edbd1bb1d3f1789b100edb79e95ab7cc4f255f0 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Tue, 26 Nov 2024 23:03:58 +0000 Subject: [PATCH 332/829] feat(Algebra/GradedMonoid): add `prod_mem_graded` (#19500) Co-authored-by: Kevin Buzzard This contribution was created as part of the Durham Computational Algebraic Geometry Workshop. --- Mathlib/Algebra/GradedMonoid.lean | 30 ++++++++++++++++++++++++++++++ 1 file changed, 30 insertions(+) diff --git a/Mathlib/Algebra/GradedMonoid.lean b/Mathlib/Algebra/GradedMonoid.lean index 4371d6d89426e..a4aa2463b4229 100644 --- a/Mathlib/Algebra/GradedMonoid.lean +++ b/Mathlib/Algebra/GradedMonoid.lean @@ -9,6 +9,7 @@ import Mathlib.Algebra.Group.Submonoid.Defs import Mathlib.Data.List.FinRange import Mathlib.Data.SetLike.Basic import Mathlib.Data.Sigma.Basic +import Mathlib.Algebra.BigOperators.Group.Finset import Lean.Elab.Tactic /-! @@ -671,3 +672,32 @@ def SetLike.homogeneousSubmonoid [AddMonoid ι] [Monoid R] (A : ι → S) [SetLi mul_mem' a b := SetLike.homogeneous_mul a b end HomogeneousElements + +section CommMonoid + +namespace SetLike + +variable {ι R S : Type*} [SetLike S R] [CommMonoid R] [AddCommMonoid ι] +variable (A : ι → S) [SetLike.GradedMonoid A] + +variable {κ : Type*} (i : κ → ι) (g : κ → R) {F : Finset κ} + +theorem prod_mem_graded (hF : ∀ k ∈ F, g k ∈ A (i k)) : ∏ k ∈ F, g k ∈ A (∑ k ∈ F, i k) := by + classical + induction F using Finset.induction_on + · simp [GradedOne.one_mem] + · case insert j F' hF2 h3 => + rw [Finset.prod_insert hF2, Finset.sum_insert hF2] + apply SetLike.mul_mem_graded (hF j <| Finset.mem_insert_self j F') + apply h3 + intro k hk + apply hF k + exact Finset.mem_insert_of_mem hk + +theorem prod_pow_mem_graded (n : κ → ℕ) (hF : ∀ k ∈ F, g k ∈ A (i k)) : + ∏ k ∈ F, g k ^ n k ∈ A (∑ k ∈ F, n k • i k) := + prod_mem_graded A _ _ fun k hk ↦ pow_mem_graded _ (hF k hk) + +end SetLike + +end CommMonoid From ba018bd9cc58c0eff74c88301e2447b01accb47e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Tue, 26 Nov 2024 23:14:43 +0000 Subject: [PATCH 333/829] refactor(CategoryTheory/Sites): remove SheafOfTypes (#19510) This PR removes the category `SheafOfTypes.{u} J` which was equivalent to `Sheaf J (Type u)`, where `J` is a Grothendieck topology. This avoids some duplication of code, as `SheafOfTypes` was essentially a special case of `Sheaf`: the only difference was that the sheaf condition was phrased differently. --- Mathlib/CategoryTheory/Sites/Adjunction.lean | 41 +++-------- .../Sites/Coherent/LocallySurjective.lean | 15 ++-- Mathlib/CategoryTheory/Sites/Continuous.lean | 24 ++++--- .../CategoryTheory/Sites/CoverPreserving.lean | 11 +-- .../CategoryTheory/Sites/DenseSubsite.lean | 24 ++++--- .../Sites/LocallyFullyFaithful.lean | 10 +-- Mathlib/CategoryTheory/Sites/Sheaf.lean | 38 +++++----- .../CategoryTheory/Sites/SheafOfTypes.lean | 70 ------------------- Mathlib/CategoryTheory/Sites/Types.lean | 52 +++++++++----- 9 files changed, 116 insertions(+), 169 deletions(-) diff --git a/Mathlib/CategoryTheory/Sites/Adjunction.lean b/Mathlib/CategoryTheory/Sites/Adjunction.lean index 01604a3fc6ead..8b4192d40635c 100644 --- a/Mathlib/CategoryTheory/Sites/Adjunction.lean +++ b/Mathlib/CategoryTheory/Sites/Adjunction.lean @@ -29,8 +29,8 @@ variable {F : D ⥤ E} {G : E ⥤ D} /-- The forgetful functor from `Sheaf J D` to sheaves of types, for a concrete category `D` whose forgetful functor preserves the correct limits. -/ abbrev sheafForget [ConcreteCategory D] [HasSheafCompose J (forget D)] : - Sheaf J D ⥤ SheafOfTypes J := - sheafCompose J (forget D) ⋙ (sheafEquivSheafOfTypes J).functor + Sheaf J D ⥤ Sheaf J (Type _) := + sheafCompose J (forget D) namespace Sheaf @@ -98,37 +98,16 @@ section ForgetToType variable [HasWeakSheafify J D] [ConcreteCategory D] [HasSheafCompose J (forget D)] -/-- This is the functor sending a sheaf of types `X` to the sheafification of `X ⋙ G`. -/ -abbrev composeAndSheafifyFromTypes (G : Type max v₁ u₁ ⥤ D) : SheafOfTypes J ⥤ Sheaf J D := - (sheafEquivSheafOfTypes J).inverse ⋙ composeAndSheafify _ G +@[deprecated (since := "2024-11-26")] alias composeAndSheafifyFromTypes := composeAndSheafify -/-- A variant of the adjunction between sheaf categories, in the case where the right adjoint -is the forgetful functor to sheaves of types. -/ -def adjunctionToTypes {G : Type max v₁ u₁ ⥤ D} (adj : G ⊣ forget D) : - composeAndSheafifyFromTypes J G ⊣ sheafForget J := - (sheafEquivSheafOfTypes J).symm.toAdjunction.comp (adjunction J adj) +/-- The adjunction `composeAndSheafify J G ⊣ sheafForget J`. (Use `Sheaf.adjunction`.)-/ +@[deprecated (since := "2024-11-26")] abbrev adjunctionToTypes + {G : Type max v₁ u₁ ⥤ D} (adj : G ⊣ forget D) : + composeAndSheafify J G ⊣ sheafForget J := + adjunction _ adj -@[simp] -theorem adjunctionToTypes_unit_app_val {G : Type max v₁ u₁ ⥤ D} (adj : G ⊣ forget D) - (Y : SheafOfTypes J) : - ((adjunctionToTypes J adj).unit.app Y).val = - (adj.whiskerRight _).unit.app ((sheafOfTypesToPresheaf J).obj Y) ≫ - whiskerRight (toSheafify J _) (forget D) := by - simp [adjunctionToTypes] - rfl - -@[simp] -theorem adjunctionToTypes_counit_app_val {G : Type max v₁ u₁ ⥤ D} (adj : G ⊣ forget D) - (X : Sheaf J D) : - ((adjunctionToTypes J adj).counit.app X).val = - sheafifyLift J ((Functor.associator _ _ _).hom ≫ (adj.whiskerRight _).counit.app _) X.2 := by - apply sheafifyLift_unique - ext - simp [adjunctionToTypes, sheafEquivSheafOfTypes, Equivalence.symm] - -instance [(forget D).IsRightAdjoint] : - (sheafForget.{_, _, _, _, max u₁ v₁} (D := D) J).IsRightAdjoint := - (adjunctionToTypes J (Adjunction.ofIsRightAdjoint (forget D))).isRightAdjoint +example [(forget D).IsRightAdjoint] : + (sheafForget.{_, _, _, _, max u₁ v₁} (D := D) J).IsRightAdjoint := by infer_instance end ForgetToType diff --git a/Mathlib/CategoryTheory/Sites/Coherent/LocallySurjective.lean b/Mathlib/CategoryTheory/Sites/Coherent/LocallySurjective.lean index bd98aeb91cd38..e7d77d4762d54 100644 --- a/Mathlib/CategoryTheory/Sites/Coherent/LocallySurjective.lean +++ b/Mathlib/CategoryTheory/Sites/Coherent/LocallySurjective.lean @@ -53,7 +53,7 @@ lemma regularTopology.isLocallySurjective_iff [Preregular C] {F G : Cᵒᵖ ⥤ rw [regularTopology.mem_sieves_iff_hasEffectiveEpi] exact ⟨X', π, h, h'⟩ -lemma extensiveTopology.surjective_of_isLocallySurjective_sheafOfTypes [FinitaryPreExtensive C] +lemma extensiveTopology.surjective_of_isLocallySurjective_sheaf_of_types [FinitaryPreExtensive C] {F G : Cᵒᵖ ⥤ Type w} (f : F ⟶ G) [PreservesFiniteProducts F] [PreservesFiniteProducts G] (h : Presheaf.IsLocallySurjective (extensiveTopology C) f) {X : C} : Function.Surjective (f.app (op X)) := by @@ -84,13 +84,17 @@ lemma extensiveTopology.surjective_of_isLocallySurjective_sheafOfTypes [Finitary erw [IsLimit.map_π] rfl +@[deprecated (since := "2024-11-26")] +alias extensiveTopology.surjective_of_isLocallySurjective_sheafOfTypes := + extensiveTopology.surjective_of_isLocallySurjective_sheaf_of_types + lemma extensiveTopology.presheafIsLocallySurjective_iff [FinitaryPreExtensive C] {F G : Cᵒᵖ ⥤ D} (f : F ⟶ G) [PreservesFiniteProducts F] [PreservesFiniteProducts G] [PreservesFiniteProducts (forget D)] : Presheaf.IsLocallySurjective (extensiveTopology C) f ↔ ∀ (X : C), Function.Surjective (f.app (op X)) := by constructor · rw [Presheaf.isLocallySurjective_iff_whisker_forget (J := extensiveTopology C)] - exact fun h _ ↦ surjective_of_isLocallySurjective_sheafOfTypes (whiskerRight f (forget D)) h + exact fun h _ ↦ surjective_of_isLocallySurjective_sheaf_of_types (whiskerRight f (forget D)) h · intro h refine ⟨fun {X} y ↦ ?_⟩ obtain ⟨x, hx⟩ := h X y @@ -104,7 +108,7 @@ lemma extensiveTopology.isLocallySurjective_iff [FinitaryExtensive C] ∀ (X : C), Function.Surjective (f.val.app (op X)) := extensiveTopology.presheafIsLocallySurjective_iff _ f.val -lemma regularTopology.isLocallySurjective_sheafOfTypes [Preregular C] [FinitaryPreExtensive C] +lemma regularTopology.isLocallySurjective_sheaf_of_types [Preregular C] [FinitaryPreExtensive C] {F G : Cᵒᵖ ⥤ Type w} (f : F ⟶ G) [PreservesFiniteProducts F] [PreservesFiniteProducts G] (h : Presheaf.IsLocallySurjective (coherentTopology C) f) : Presheaf.IsLocallySurjective (regularTopology C) f where @@ -136,6 +140,9 @@ lemma regularTopology.isLocallySurjective_sheafOfTypes [Preregular C] [FinitaryP · change G.map _ (G.map _ _) = _ simp only [← FunctorToTypes.map_comp_apply, ← op_comp, Sigma.ι_desc] +@[deprecated (since := "2024-11-26")] alias regularTopology.isLocallySurjective_sheafOfTypes := +regularTopology.isLocallySurjective_sheaf_of_types + lemma coherentTopology.presheafIsLocallySurjective_iff {F G : Cᵒᵖ ⥤ D} (f : F ⟶ G) [Preregular C] [FinitaryPreExtensive C] [PreservesFiniteProducts F] [PreservesFiniteProducts G] [PreservesFiniteProducts (forget D)] : @@ -144,7 +151,7 @@ lemma coherentTopology.presheafIsLocallySurjective_iff {F G : Cᵒᵖ ⥤ D} (f constructor · rw [Presheaf.isLocallySurjective_iff_whisker_forget, Presheaf.isLocallySurjective_iff_whisker_forget (J := regularTopology C)] - exact regularTopology.isLocallySurjective_sheafOfTypes _ + exact regularTopology.isLocallySurjective_sheaf_of_types _ · refine Presheaf.isLocallySurjective_of_le (J := regularTopology C) ?_ _ rw [← extensive_regular_generate_coherent] exact (Coverage.gi _).gc.monotone_l le_sup_right diff --git a/Mathlib/CategoryTheory/Sites/Continuous.lean b/Mathlib/CategoryTheory/Sites/Continuous.lean index 91469d869e53b..ef0c9ef232567 100644 --- a/Mathlib/CategoryTheory/Sites/Continuous.lean +++ b/Mathlib/CategoryTheory/Sites/Continuous.lean @@ -115,31 +115,35 @@ abbrev PreservesOneHypercovers := /-- A functor `F` is continuous if the precomposition with `F.op` sends sheaves of `Type t` to sheaves. -/ class IsContinuous : Prop where - op_comp_isSheafOfTypes (G : SheafOfTypes.{t} K) : Presieve.IsSheaf J (F.op ⋙ G.val) + op_comp_isSheaf_of_types (G : Sheaf K (Type t)) : Presieve.IsSheaf J (F.op ⋙ G.val) -lemma op_comp_isSheafOfTypes [Functor.IsContinuous.{t} F J K] (G : SheafOfTypes.{t} K) : +lemma op_comp_isSheaf_of_types [Functor.IsContinuous.{t} F J K] (G : Sheaf K (Type t)) : Presieve.IsSheaf J (F.op ⋙ G.val) := - Functor.IsContinuous.op_comp_isSheafOfTypes _ + Functor.IsContinuous.op_comp_isSheaf_of_types _ + +@[deprecated (since := "2024-11-26")] alias op_comp_isSheafOfTypes := op_comp_isSheaf_of_types lemma op_comp_isSheaf [Functor.IsContinuous.{t} F J K] (G : Sheaf K A) : Presheaf.IsSheaf J (F.op ⋙ G.val) := - fun T => F.op_comp_isSheafOfTypes J K ⟨_, G.cond T⟩ + fun T => F.op_comp_isSheaf_of_types J K ⟨_, (isSheaf_iff_isSheaf_of_type _ _).2 (G.cond T)⟩ lemma isContinuous_of_iso {F₁ F₂ : C ⥤ D} (e : F₁ ≅ F₂) (J : GrothendieckTopology C) (K : GrothendieckTopology D) [Functor.IsContinuous.{t} F₁ J K] : Functor.IsContinuous.{t} F₂ J K where - op_comp_isSheafOfTypes G := + op_comp_isSheaf_of_types G := Presieve.isSheaf_iso J (isoWhiskerRight (NatIso.op e.symm) _) - (F₁.op_comp_isSheafOfTypes J K G) + (F₁.op_comp_isSheaf_of_types J K G) instance isContinuous_id : Functor.IsContinuous.{w} (𝟭 C) J J where - op_comp_isSheafOfTypes G := G.2 + op_comp_isSheaf_of_types G := (isSheaf_iff_isSheaf_of_type _ _).1 G.2 lemma isContinuous_comp (F₁ : C ⥤ D) (F₂ : D ⥤ E) (J : GrothendieckTopology C) (K : GrothendieckTopology D) (L : GrothendieckTopology E) [Functor.IsContinuous.{t} F₁ J K] [Functor.IsContinuous.{t} F₂ K L] : Functor.IsContinuous.{t} (F₁ ⋙ F₂) J L where - op_comp_isSheafOfTypes G := F₁.op_comp_isSheafOfTypes J K ⟨_, F₂.op_comp_isSheafOfTypes K L G⟩ + op_comp_isSheaf_of_types G := + F₁.op_comp_isSheaf_of_types J K + ⟨_,(isSheaf_iff_isSheaf_of_type _ _).2 (F₂.op_comp_isSheaf_of_types K L G)⟩ lemma isContinuous_comp' {F₁ : C ⥤ D} {F₂ : D ⥤ E} {F₁₂ : C ⥤ E} (e : F₁ ⋙ F₂ ≅ F₁₂) (J : GrothendieckTopology C) @@ -163,9 +167,9 @@ lemma op_comp_isSheaf_of_preservesOneHypercovers lemma isContinuous_of_preservesOneHypercovers [PreservesOneHypercovers.{w} F J K] [GrothendieckTopology.IsGeneratedByOneHypercovers.{w} J] : IsContinuous.{t} F J K where - op_comp_isSheafOfTypes := by + op_comp_isSheaf_of_types := by rintro ⟨P, hP⟩ - rw [← isSheaf_iff_isSheaf_of_type] at hP ⊢ + rw [← isSheaf_iff_isSheaf_of_type] exact F.op_comp_isSheaf_of_preservesOneHypercovers J K P hP end diff --git a/Mathlib/CategoryTheory/Sites/CoverPreserving.lean b/Mathlib/CategoryTheory/Sites/CoverPreserving.lean index 54301e987ef0b..67fded0c9b4c2 100644 --- a/Mathlib/CategoryTheory/Sites/CoverPreserving.lean +++ b/Mathlib/CategoryTheory/Sites/CoverPreserving.lean @@ -75,13 +75,13 @@ This is actually stronger than merely preserving compatible families because of -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not ported yet @[nolint has_nonempty_instance] structure CompatiblePreserving (K : GrothendieckTopology D) (G : C ⥤ D) : Prop where compatible : - ∀ (ℱ : SheafOfTypes.{w} K) {Z} {T : Presieve Z} {x : FamilyOfElements (G.op ⋙ ℱ.val) T} + ∀ (ℱ : Sheaf K (Type w)) {Z} {T : Presieve Z} {x : FamilyOfElements (G.op ⋙ ℱ.val) T} (_ : x.Compatible) {Y₁ Y₂} {X} (f₁ : X ⟶ G.obj Y₁) (f₂ : X ⟶ G.obj Y₂) {g₁ : Y₁ ⟶ Z} {g₂ : Y₂ ⟶ Z} (hg₁ : T g₁) (hg₂ : T g₂) (_ : f₁ ≫ G.map g₁ = f₂ ≫ G.map g₂), ℱ.val.map f₁.op (x g₁ hg₁) = ℱ.val.map f₂.op (x g₂ hg₂) section -variable {J K} {G : C ⥤ D} (hG : CompatiblePreserving.{w} K G) (ℱ : SheafOfTypes.{w} K) {Z : C} +variable {J K} {G : C ⥤ D} (hG : CompatiblePreserving.{w} K G) (ℱ : Sheaf K (Type w)) {Z : C} variable {T : Presieve Z} {x : FamilyOfElements (G.op ⋙ ℱ.val) T} (h : x.Compatible) include hG h @@ -163,14 +163,15 @@ This result is basically . -/ lemma Functor.isContinuous_of_coverPreserving (hF₁ : CompatiblePreserving.{w} K F) (hF₂ : CoverPreserving J K F) : Functor.IsContinuous.{w} F J K where - op_comp_isSheafOfTypes G X S hS x hx := by + op_comp_isSheaf_of_types G X S hS x hx := by apply exists_unique_of_exists_of_unique - · have H := G.2 _ (hF₂.cover_preserve hS) + · have H := (isSheaf_iff_isSheaf_of_type _ _).1 G.2 _ (hF₂.cover_preserve hS) exact ⟨H.amalgamate (x.functorPushforward F) (hx.functorPushforward hF₁), fun V f hf => (H.isAmalgamation (hx.functorPushforward hF₁) (F.map f) _).trans (hF₁.apply_map _ hx hf)⟩ · intro y₁ y₂ hy₁ hy₂ - apply (Presieve.isSeparated_of_isSheaf _ _ G.cond _ (hF₂.cover_preserve hS)).ext + apply (Presieve.isSeparated_of_isSheaf _ _ ((isSheaf_iff_isSheaf_of_type _ _).1 G.2) _ + (hF₂.cover_preserve hS)).ext rintro Y _ ⟨Z, g, h, hg, rfl⟩ dsimp simp only [Functor.map_comp, types_comp_apply] diff --git a/Mathlib/CategoryTheory/Sites/DenseSubsite.lean b/Mathlib/CategoryTheory/Sites/DenseSubsite.lean index fb9eb889b3ee1..5e24ab6fe6ff6 100644 --- a/Mathlib/CategoryTheory/Sites/DenseSubsite.lean +++ b/Mathlib/CategoryTheory/Sites/DenseSubsite.lean @@ -115,9 +115,10 @@ variable {K} variable {A : Type*} [Category A] (G : C ⥤ D) -- this is not marked with `@[ext]` because `H` can not be inferred from the type -theorem ext [G.IsCoverDense K] (ℱ : SheafOfTypes K) (X : D) {s t : ℱ.val.obj (op X)} +theorem ext [G.IsCoverDense K] (ℱ : Sheaf K (Type _)) (X : D) {s t : ℱ.val.obj (op X)} (h : ∀ ⦃Y : C⦄ (f : G.obj Y ⟶ X), ℱ.val.map f.op s = ℱ.val.map f.op t) : s = t := by - apply (ℱ.cond (Sieve.coverByImage G X) (G.is_cover_of_isCoverDense K X)).isSeparatedFor.ext + apply ((isSheaf_iff_isSheaf_of_type _ _ ).1 ℱ.cond + (Sieve.coverByImage G X) (G.is_cover_of_isCoverDense K X)).isSeparatedFor.ext rintro Y _ ⟨Z, f₁, f₂, ⟨rfl⟩⟩ simp [h f₂] @@ -156,7 +157,7 @@ theorem sheaf_eq_amalgamation (ℱ : Sheaf K A) {X : A} {U : D} {T : Sieve U} (h namespace Types -variable {ℱ : Dᵒᵖ ⥤ Type v} {ℱ' : SheafOfTypes.{v} K} (α : G.op ⋙ ℱ ⟶ G.op ⋙ ℱ'.val) +variable {ℱ : Dᵒᵖ ⥤ Type v} {ℱ' : Sheaf K (Type v)} (α : G.op ⋙ ℱ ⟶ G.op ⋙ ℱ'.val) theorem naturality_apply [G.IsLocallyFull K] {X Y : C} (i : G.obj X ⟶ G.obj Y) (x) : ℱ'.1.map i.op (α.app _ x) = α.app _ (ℱ.map i.op x) := by @@ -212,13 +213,14 @@ theorem pushforwardFamily_compatible {X} (x : ℱ.obj (op X)) : /-- (Implementation). The morphism `ℱ(X) ⟶ ℱ'(X)` given by gluing the `pushforwardFamily`. -/ noncomputable def appHom (X : D) : ℱ.obj (op X) ⟶ ℱ'.val.obj (op X) := fun x => - (ℱ'.cond _ (G.is_cover_of_isCoverDense _ X)).amalgamate (pushforwardFamily α x) - (pushforwardFamily_compatible α x) + ((isSheaf_iff_isSheaf_of_type _ _ ).1 ℱ'.cond _ + (G.is_cover_of_isCoverDense _ X)).amalgamate (pushforwardFamily α x) + (pushforwardFamily_compatible α x) @[simp] theorem appHom_restrict {X : D} {Y : C} (f : op X ⟶ op (G.obj Y)) (x) : ℱ'.val.map f (appHom α X x) = α.app (op Y) (ℱ.map f x) := - ((ℱ'.cond _ (G.is_cover_of_isCoverDense _ X)).valid_glue + (((isSheaf_iff_isSheaf_of_type _ _ ).1 ℱ'.cond _ (G.is_cover_of_isCoverDense _ X)).valid_glue (pushforwardFamily_compatible α x) f.unop (Presieve.in_coverByImage G f.unop)).trans (pushforwardFamily_apply _ _ _) @@ -232,7 +234,7 @@ theorem appHom_valid_glue {X : D} {Y : C} (f : op X ⟶ op (G.obj Y)) : (Implementation). The maps given in `appIso` is inverse to each other and gives a `ℱ(X) ≅ ℱ'(X)`. -/ @[simps] -noncomputable def appIso {ℱ ℱ' : SheafOfTypes.{v} K} (i : G.op ⋙ ℱ.val ≅ G.op ⋙ ℱ'.val) +noncomputable def appIso {ℱ ℱ' : Sheaf K (Type v)} (i : G.op ⋙ ℱ.val ≅ G.op ⋙ ℱ'.val) (X : D) : ℱ.val.obj (op X) ≅ ℱ'.val.obj (op X) where hom := appHom i.hom X inv := appHom i.inv X @@ -267,7 +269,7 @@ where `G` is locally-full and cover-dense, and `ℱ, ℱ'` are sheaves, we may obtain a natural isomorphism between presheaves. -/ @[simps!] -noncomputable def presheafIso {ℱ ℱ' : SheafOfTypes.{v} K} (i : G.op ⋙ ℱ.val ≅ G.op ⋙ ℱ'.val) : +noncomputable def presheafIso {ℱ ℱ' : Sheaf K (Type v)} (i : G.op ⋙ ℱ.val ≅ G.op ⋙ ℱ'.val) : ℱ.val ≅ ℱ'.val := NatIso.ofComponents (fun X => appIso i (unop X)) @(presheafHom i.hom).naturality @@ -277,7 +279,7 @@ where `G` is locally-full and cover-dense, and `ℱ, ℱ'` are sheaves, we may obtain a natural isomorphism between sheaves. -/ @[simps] -noncomputable def sheafIso {ℱ ℱ' : SheafOfTypes.{v} K} (i : G.op ⋙ ℱ.val ≅ G.op ⋙ ℱ'.val) : +noncomputable def sheafIso {ℱ ℱ' : Sheaf K (Type v)} (i : G.op ⋙ ℱ.val ≅ G.op ⋙ ℱ'.val) : ℱ ≅ ℱ' where hom := ⟨(presheafIso i).hom⟩ inv := ⟨(presheafIso i).inv⟩ @@ -398,7 +400,9 @@ theorem sheafHom_restrict_eq (α : G.op ⋙ ℱ ⟶ G.op ⋙ ℱ'.val) : simp only [Category.assoc] congr 1 have := naturality_apply (G := G) (ℱ := ℱ ⋙ coyoneda.obj (op <| (G.op ⋙ ℱ).obj X)) - (ℱ' := ⟨_, ℱ'.2 ((G.op ⋙ ℱ).obj X)⟩) (whiskerRight α (coyoneda.obj _)) hf.some.map (𝟙 _) + (ℱ' := ⟨_, Presheaf.isSheaf_comp_of_isSheaf K ℱ'.val + (coyoneda.obj (op ((G.op ⋙ ℱ).obj X))) ℱ'.cond⟩) + (whiskerRight α (coyoneda.obj _)) hf.some.map (𝟙 _) simpa using this variable (G) diff --git a/Mathlib/CategoryTheory/Sites/LocallyFullyFaithful.lean b/Mathlib/CategoryTheory/Sites/LocallyFullyFaithful.lean index 96bf30d33b13f..2988aa4190f18 100644 --- a/Mathlib/CategoryTheory/Sites/LocallyFullyFaithful.lean +++ b/Mathlib/CategoryTheory/Sites/LocallyFullyFaithful.lean @@ -103,20 +103,22 @@ variable {K} variable {A : Type*} [Category A] (G : C ⥤ D) theorem IsLocallyFull.ext [G.IsLocallyFull K] - (ℱ : SheafOfTypes K) {X Y : C} (i : G.obj X ⟶ G.obj Y) + (ℱ : Sheaf K (Type _)) {X Y : C} (i : G.obj X ⟶ G.obj Y) {s t : ℱ.val.obj (op (G.obj X))} (h : ∀ ⦃Z : C⦄ (j : Z ⟶ X) (f : Z ⟶ Y), G.map f = G.map j ≫ i → ℱ.1.map (G.map j).op s = ℱ.1.map (G.map j).op t) : s = t := by - apply (ℱ.cond _ (G.functorPushforward_imageSieve_mem K i)).isSeparatedFor.ext + apply (((isSheaf_iff_isSheaf_of_type _ _).1 ℱ.cond) _ + (G.functorPushforward_imageSieve_mem K i)).isSeparatedFor.ext rintro Z _ ⟨W, iWX, iZW, ⟨iWY, e⟩, rfl⟩ simp [h iWX iWY e] -theorem IsLocallyFaithful.ext [G.IsLocallyFaithful K] (ℱ : SheafOfTypes K) +theorem IsLocallyFaithful.ext [G.IsLocallyFaithful K] (ℱ : Sheaf K (Type _)) {X Y : C} (i₁ i₂ : X ⟶ Y) (e : G.map i₁ = G.map i₂) {s t : ℱ.val.obj (op (G.obj X))} (h : ∀ ⦃Z : C⦄ (j : Z ⟶ X), j ≫ i₁ = j ≫ i₂ → ℱ.1.map (G.map j).op s = ℱ.1.map (G.map j).op t) : s = t := by - apply (ℱ.cond _ (G.functorPushforward_equalizer_mem K i₁ i₂ e)).isSeparatedFor.ext + apply (((isSheaf_iff_isSheaf_of_type _ _).1 ℱ.cond) _ + (G.functorPushforward_equalizer_mem K i₁ i₂ e)).isSeparatedFor.ext rintro Z _ ⟨W, iWX, iZW, hiWX, rfl⟩ simp [h iWX hiWX] diff --git a/Mathlib/CategoryTheory/Sites/Sheaf.lean b/Mathlib/CategoryTheory/Sites/Sheaf.lean index 9ef621b55f223..6900396ec52a3 100644 --- a/Mathlib/CategoryTheory/Sites/Sheaf.lean +++ b/Mathlib/CategoryTheory/Sites/Sheaf.lean @@ -376,12 +376,6 @@ theorem Sheaf.Hom.mono_of_presheaf_mono {F G : Sheaf J A} (f : F ⟶ G) [h : Mon instance Sheaf.Hom.epi_of_presheaf_epi {F G : Sheaf J A} (f : F ⟶ G) [h : Epi f.1] : Epi f := (sheafToPresheaf J A).epi_of_epi_map h -/-- The sheaf of sections guaranteed by the sheaf condition. -/ -@[simps] -def sheafOver {A : Type u₂} [Category.{v₂} A] {J : GrothendieckTopology C} (ℱ : Sheaf J A) (E : A) : - SheafOfTypes J := - ⟨ℱ.val ⋙ coyoneda.obj (op E), ℱ.cond E⟩ - theorem isSheaf_iff_isSheaf_of_type (P : Cᵒᵖ ⥤ Type w) : Presheaf.IsSheaf J P ↔ Presieve.IsSheaf J P := by constructor @@ -402,27 +396,39 @@ theorem isSheaf_iff_isSheaf_of_type (P : Cᵒᵖ ⥤ Type w) : rw [Presieve.IsSheafFor.valid_glue _ _ _ hf, ← hy _ hf] rfl +/-- The sheaf of sections guaranteed by the sheaf condition. -/ +@[simps] +def sheafOver {A : Type u₂} [Category.{v₂} A] {J : GrothendieckTopology C} (ℱ : Sheaf J A) (E : A) : + Sheaf J (Type _) where + val := ℱ.val ⋙ coyoneda.obj (op E) + cond := by + rw [isSheaf_iff_isSheaf_of_type] + exact ℱ.cond E + variable {J} in lemma Presheaf.IsSheaf.isSheafFor {P : Cᵒᵖ ⥤ Type w} (hP : Presheaf.IsSheaf J P) {X : C} (S : Sieve X) (hS : S ∈ J X) : Presieve.IsSheafFor P S.arrows := by rw [isSheaf_iff_isSheaf_of_type] at hP exact hP S hS -/-- The category of sheaves taking values in Type is the same as the category of set-valued sheaves. +variable {A} in +lemma Presheaf.isSheaf_bot (P : Cᵒᵖ ⥤ A) : IsSheaf ⊥ P := fun _ ↦ Presieve.isSheaf_bot + +/-- +The category of sheaves on the bottom (trivial) Grothendieck topology is +equivalent to the category of presheaves. -/ @[simps] -def sheafEquivSheafOfTypes : Sheaf J (Type w) ≌ SheafOfTypes J where - functor := - { obj := fun S => ⟨S.val, (isSheaf_iff_isSheaf_of_type _ _).1 S.2⟩ - map := fun f => ⟨f.val⟩ } +def sheafBotEquivalence : Sheaf (⊥ : GrothendieckTopology C) A ≌ Cᵒᵖ ⥤ A where + functor := sheafToPresheaf _ _ inverse := - { obj := fun S => ⟨S.val, (isSheaf_iff_isSheaf_of_type _ _).2 S.2⟩ - map := fun f => ⟨f.val⟩ } - unitIso := NatIso.ofComponents fun _ => Iso.refl _ - counitIso := NatIso.ofComponents fun _ => Iso.refl _ + { obj := fun P => ⟨P, Presheaf.isSheaf_bot P⟩ + map := fun f => ⟨f⟩ } + unitIso := Iso.refl _ + counitIso := Iso.refl _ instance : Inhabited (Sheaf (⊥ : GrothendieckTopology C) (Type w)) := - ⟨(sheafEquivSheafOfTypes _).inverse.obj default⟩ + ⟨(sheafBotEquivalence _).inverse.obj ((Functor.const _).obj default)⟩ variable {J} {A} diff --git a/Mathlib/CategoryTheory/Sites/SheafOfTypes.lean b/Mathlib/CategoryTheory/Sites/SheafOfTypes.lean index ec2ba32663333..9819869ea4a96 100644 --- a/Mathlib/CategoryTheory/Sites/SheafOfTypes.lean +++ b/Mathlib/CategoryTheory/Sites/SheafOfTypes.lean @@ -198,74 +198,4 @@ theorem forallYonedaIsSheaf_iff_colimit (S : Sieve X) : end Sieve -variable {C : Type u} [Category.{v} C] -variable (J : GrothendieckTopology C) - -/-- The category of sheaves on a grothendieck topology. -/ -structure SheafOfTypes (J : GrothendieckTopology C) : Type max u v (w + 1) where - /-- the underlying presheaf -/ - val : Cᵒᵖ ⥤ Type w - /-- the condition that the presheaf is a sheaf -/ - cond : Presieve.IsSheaf J val - -namespace SheafOfTypes - -variable {J} - -/-- Morphisms between sheaves of types are just morphisms between the underlying presheaves. -/ -@[ext] -structure Hom (X Y : SheafOfTypes J) where - /-- a morphism between the underlying presheaves -/ - val : X.val ⟶ Y.val - -@[simps] -instance : Category (SheafOfTypes J) where - Hom := Hom - id _ := ⟨𝟙 _⟩ - comp f g := ⟨f.val ≫ g.val⟩ - id_comp _ := Hom.ext <| id_comp _ - comp_id _ := Hom.ext <| comp_id _ - assoc _ _ _ := Hom.ext <| assoc _ _ _ - --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): we need to restate the `ext` lemma in terms of the categorical morphism. --- not just the underlying structure. --- It would be nice if this boilerplate weren't necessary. -@[ext] -theorem Hom.ext' {X Y : SheafOfTypes J} (f g : X ⟶ Y) (w : f.val = g.val) : f = g := - Hom.ext w - --- Let's make the inhabited linter happy... -instance (X : SheafOfTypes J) : Inhabited (Hom X X) := - ⟨𝟙 X⟩ - -end SheafOfTypes - -/-- The inclusion functor from sheaves to presheaves. -/ -@[simps] -def sheafOfTypesToPresheaf : SheafOfTypes J ⥤ Cᵒᵖ ⥤ Type w where - obj := SheafOfTypes.val - map f := f.val - map_id _ := rfl - map_comp _ _ := rfl - -instance : (sheafOfTypesToPresheaf J).Full where map_surjective f := ⟨⟨f⟩, rfl⟩ - -instance : (sheafOfTypesToPresheaf J).Faithful where - -/-- -The category of sheaves on the bottom (trivial) grothendieck topology is equivalent to the category -of presheaves. --/ -@[simps] -def sheafOfTypesBotEquiv : SheafOfTypes (⊥ : GrothendieckTopology C) ≌ Cᵒᵖ ⥤ Type w where - functor := sheafOfTypesToPresheaf _ - inverse := - { obj := fun P => ⟨P, Presieve.isSheaf_bot⟩ - map := fun f => ⟨f⟩ } - unitIso := Iso.refl _ - counitIso := Iso.refl _ - -instance : Inhabited (SheafOfTypes (⊥ : GrothendieckTopology C)) := - ⟨sheafOfTypesBotEquiv.inverse.obj ((Functor.const _).obj PUnit)⟩ - end CategoryTheory diff --git a/Mathlib/CategoryTheory/Sites/Types.lean b/Mathlib/CategoryTheory/Sites/Types.lean index cfa559d44c11a..65430a4c42150 100644 --- a/Mathlib/CategoryTheory/Sites/Types.lean +++ b/Mathlib/CategoryTheory/Sites/Types.lean @@ -45,23 +45,31 @@ theorem generate_discretePresieve_mem (α : Type u) : Sieve.generate (discretePresieve α) ∈ typesGrothendieckTopology α := fun x => ⟨PUnit, id, fun _ => x, ⟨PUnit.unit, fun _ => Subsingleton.elim _ _⟩, rfl⟩ -open Presieve - -theorem isSheaf_yoneda' {α : Type u} : IsSheaf typesGrothendieckTopology (yoneda.obj α) := +/-- The sheaf condition for `yoneda'`. -/ +theorem Presieve.isSheaf_yoneda' {α : Type u} : + Presieve.IsSheaf typesGrothendieckTopology (yoneda.obj α) := fun β _ hs x hx => ⟨fun y => x _ (hs y) PUnit.unit, fun γ f h => funext fun z => by convert congr_fun (hx (𝟙 _) (fun _ => z) (hs <| f z) h rfl) PUnit.unit using 1, fun f hf => funext fun y => by convert congr_fun (hf _ (hs y)) PUnit.unit⟩ +/-- The sheaf condition for `yoneda'`. -/ +theorem Presheaf.isSheaf_yoneda' {α : Type u} : + Presheaf.IsSheaf typesGrothendieckTopology (yoneda.obj α) := by + rw [isSheaf_iff_isSheaf_of_type] + exact Presieve.isSheaf_yoneda' + +@[deprecated (since := "2024-11-26")] alias isSheaf_yoneda' := Presieve.isSheaf_yoneda' + /-- The yoneda functor that sends a type to a sheaf over the category of types. -/ @[simps] -def yoneda' : Type u ⥤ SheafOfTypes typesGrothendieckTopology where - obj α := ⟨yoneda.obj α, isSheaf_yoneda'⟩ +def yoneda' : Type u ⥤ Sheaf typesGrothendieckTopology (Type u) where + obj α := ⟨yoneda.obj α, Presheaf.isSheaf_yoneda'⟩ map f := ⟨yoneda.map f⟩ @[simp] -theorem yoneda'_comp : yoneda'.{u} ⋙ sheafOfTypesToPresheaf _ = yoneda := +theorem yoneda'_comp : yoneda'.{u} ⋙ sheafToPresheaf _ _ = yoneda := rfl open Opposite @@ -71,6 +79,8 @@ a map `P(α) → (α → P(*))` for all type `α`. -/ def eval (P : Type uᵒᵖ ⥤ Type u) (α : Type u) (s : P.obj (op α)) (x : α) : P.obj (op PUnit) := P.map (↾fun _ => x).op s +open Presieve + /-- Given a sheaf `S` on the category of types, construct a map `(α → S(*)) → S(α)` that is inverse to `eval`. -/ noncomputable def typesGlue (S : Type uᵒᵖ ⥤ Type u) (hs : IsSheaf typesGrothendieckTopology S) @@ -103,10 +113,11 @@ theorem typesGlue_eval {S hs α} (s) : typesGlue.{u} S hs α (eval S α s) = s : /-- Given a sheaf `S`, construct an equivalence `S(α) ≃ (α → S(*))`. -/ @[simps] -noncomputable def evalEquiv (S : Type uᵒᵖ ⥤ Type u) (hs : IsSheaf typesGrothendieckTopology S) +noncomputable def evalEquiv (S : Type uᵒᵖ ⥤ Type u) + (hs : Presheaf.IsSheaf typesGrothendieckTopology S) (α : Type u) : S.obj (op α) ≃ (α → S.obj (op PUnit)) where toFun := eval S α - invFun := typesGlue S hs α + invFun := typesGlue S ((isSheaf_iff_isSheaf_of_type _ _ ).1 hs) α left_inv := typesGlue_eval right_inv := eval_typesGlue @@ -116,31 +127,32 @@ theorem eval_map (S : Type uᵒᵖ ⥤ Type u) (α β) (f : β ⟶ α) (s x) : /-- Given a sheaf `S`, construct an isomorphism `S ≅ [-, S(*)]`. -/ @[simps!] -noncomputable def equivYoneda (S : Type uᵒᵖ ⥤ Type u) (hs : IsSheaf typesGrothendieckTopology S) : +noncomputable def equivYoneda (S : Type uᵒᵖ ⥤ Type u) + (hs : Presheaf.IsSheaf typesGrothendieckTopology S) : S ≅ yoneda.obj (S.obj (op PUnit)) := NatIso.ofComponents (fun α => Equiv.toIso <| evalEquiv S hs <| unop α) fun {α β} f => funext fun _ => funext fun _ => eval_map S (unop α) (unop β) f.unop _ _ /-- Given a sheaf `S`, construct an isomorphism `S ≅ [-, S(*)]`. -/ @[simps] -noncomputable def equivYoneda' (S : SheafOfTypes typesGrothendieckTopology) : +noncomputable def equivYoneda' (S : Sheaf typesGrothendieckTopology (Type u)) : S ≅ yoneda'.obj (S.1.obj (op PUnit)) where hom := ⟨(equivYoneda S.1 S.2).hom⟩ inv := ⟨(equivYoneda S.1 S.2).inv⟩ hom_inv_id := by ext1; apply (equivYoneda S.1 S.2).hom_inv_id inv_hom_id := by ext1; apply (equivYoneda S.1 S.2).inv_hom_id -theorem eval_app (S₁ S₂ : SheafOfTypes.{u} typesGrothendieckTopology) (f : S₁ ⟶ S₂) (α : Type u) +theorem eval_app (S₁ S₂ : Sheaf typesGrothendieckTopology (Type u)) (f : S₁ ⟶ S₂) (α : Type u) (s : S₁.1.obj (op α)) (x : α) : eval S₂.1 α (f.val.app (op α) s) x = f.val.app (op PUnit) (eval S₁.1 α s x) := (congr_fun (f.val.naturality (↾fun _ : PUnit => x).op) s).symm /-- `yoneda'` induces an equivalence of category between `Type u` and -`SheafOfTypes typesGrothendieckTopology`. -/ +`Sheaf typesGrothendieckTopology (Type u)`. -/ @[simps!] -noncomputable def typeEquiv : Type u ≌ SheafOfTypes typesGrothendieckTopology where +noncomputable def typeEquiv : Type u ≌ Sheaf typesGrothendieckTopology (Type u) where functor := yoneda' - inverse := sheafOfTypesToPresheaf _ ⋙ (evaluation _ _).obj (op PUnit) + inverse := sheafToPresheaf _ _ ⋙ (evaluation _ _).obj (op PUnit) unitIso := NatIso.ofComponents (fun _α => -- α ≅ PUnit ⟶ α { hom := fun x _ => x @@ -149,9 +161,11 @@ noncomputable def typeEquiv : Type u ≌ SheafOfTypes typesGrothendieckTopology inv_hom_id := funext fun _ => funext fun y => PUnit.casesOn y rfl }) fun _ => rfl counitIso := Iso.symm <| - NatIso.ofComponents (fun S => equivYoneda' S) fun {S₁ S₂} f => - SheafOfTypes.Hom.ext <| NatTrans.ext <| - funext fun α => funext fun s => funext fun x => eval_app S₁ S₂ f (unop α) s x + NatIso.ofComponents (fun S => equivYoneda' S) (fun {S₁ S₂} f => by + ext ⟨α⟩ s + dsimp at s ⊢ + ext x + exact eval_app S₁ S₂ f α s x) functor_unitIso_comp X := by ext1 apply yonedaEquiv.injective @@ -159,7 +173,7 @@ noncomputable def typeEquiv : Type u ≌ SheafOfTypes typesGrothendieckTopology erw [typesGlue_eval] instance subcanonical_typesGrothendieckTopology : typesGrothendieckTopology.{u}.Subcanonical := - GrothendieckTopology.Subcanonical.of_isSheaf_yoneda_obj _ fun _ => isSheaf_yoneda' + GrothendieckTopology.Subcanonical.of_isSheaf_yoneda_obj _ fun _ => Presieve.isSheaf_yoneda' theorem typesGrothendieckTopology_eq_canonical : typesGrothendieckTopology.{u} = Sheaf.canonicalTopology (Type u) := by @@ -167,7 +181,7 @@ theorem typesGrothendieckTopology_eq_canonical : refine ⟨yoneda.obj (ULift Bool), ⟨_, rfl⟩, GrothendieckTopology.ext ?_⟩ funext α ext S - refine ⟨fun hs x => ?_, fun hs β f => isSheaf_yoneda' _ fun y => hs _⟩ + refine ⟨fun hs x => ?_, fun hs β f => Presieve.isSheaf_yoneda' _ fun y => hs _⟩ by_contra hsx have : (fun _ => ULift.up true) = fun _ => ULift.up false := (hs PUnit fun _ => x).isSeparatedFor.ext From ec29ecebbee207d4d121b387652d5b846384dc56 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Tue, 26 Nov 2024 23:28:21 +0000 Subject: [PATCH 334/829] feat(Topology/Algebra): add `IsModuleTopology.continuous_of_ringHom` (#19515) --- Mathlib/Topology/Algebra/Module/ModuleTopology.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/Mathlib/Topology/Algebra/Module/ModuleTopology.lean b/Mathlib/Topology/Algebra/Module/ModuleTopology.lean index 4de1767692be2..e5366ba9ea100 100644 --- a/Mathlib/Topology/Algebra/Module/ModuleTopology.lean +++ b/Mathlib/Topology/Algebra/Module/ModuleTopology.lean @@ -292,6 +292,16 @@ theorem continuous_neg (C : Type*) [AddCommGroup C] [Module R C] [TopologicalSpa haveI : ContinuousAdd C := IsModuleTopology.toContinuousAdd R C continuous_of_linearMap (LinearEquiv.neg R).toLinearMap +@[fun_prop, continuity] +theorem continuous_of_ringHom {R A B} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] + [TopologicalSpace R] [TopologicalSpace A] [IsModuleTopology R A] [TopologicalSpace B] + [TopologicalSemiring B] + (φ : A →+* B) (hφ : Continuous (φ.comp (algebraMap R A))) : Continuous φ := by + let inst := Module.compHom B (φ.comp (algebraMap R A)) + let φ' : A →ₗ[R] B := ⟨φ, fun r m ↦ by simp [Algebra.smul_def]; rfl⟩ + have : ContinuousSMul R B := ⟨(hφ.comp continuous_fst).mul continuous_snd⟩ + exact continuous_of_linearMap φ' + end function end IsModuleTopology From 69959750df8adb05e7a31e9aa0a8fbd850188956 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Wed, 27 Nov 2024 02:02:50 +0000 Subject: [PATCH 335/829] feat: add `NonUnital{Star}Algebra.elemental` subalgebras (#18615) Define the `NonUnital{Star}Algebra.elemental` as `(adjoin R {x}).topologicalClosure` to match the unital versions. - [ ] depends on: #18612 - [ ] depends on: #18613 --- .../Topology/Algebra/NonUnitalAlgebra.lean | 61 +++++++++++++++++ .../Algebra/NonUnitalStarAlgebra.lean | 68 +++++++++++++++++++ 2 files changed, 129 insertions(+) diff --git a/Mathlib/Topology/Algebra/NonUnitalAlgebra.lean b/Mathlib/Topology/Algebra/NonUnitalAlgebra.lean index 91c54d048c2f2..e450112e92720 100644 --- a/Mathlib/Topology/Algebra/NonUnitalAlgebra.lean +++ b/Mathlib/Topology/Algebra/NonUnitalAlgebra.lean @@ -77,3 +77,64 @@ abbrev nonUnitalCommRingTopologicalClosure [T2Space A] (s : NonUnitalSubalgebra end Ring end NonUnitalSubalgebra + +namespace NonUnitalAlgebra + +open NonUnitalSubalgebra + +variable (R : Type*) {A : Type*} [CommSemiring R] [NonUnitalSemiring A] +variable [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] +variable [TopologicalSpace A] [TopologicalSemiring A] [ContinuousConstSMul R A] + +/-- The topological closure of the non-unital subalgebra generated by a single element. -/ +def elemental (x : A) : NonUnitalSubalgebra R A := + adjoin R {x} |>.topologicalClosure + +namespace elemental + +@[aesop safe apply (rule_sets := [SetLike])] +theorem self_mem (x : A) : x ∈ elemental R x := + le_topologicalClosure _ <| self_mem_adjoin_singleton R x + +variable {R} in +theorem le_of_mem {x : A} {s : NonUnitalSubalgebra R A} (hs : IsClosed (s : Set A)) (hx : x ∈ s) : + elemental R x ≤ s := + topologicalClosure_minimal _ (adjoin_le <| by simpa using hx) hs + +variable {R} in +theorem le_iff_mem {x : A} {s : NonUnitalSubalgebra R A} (hs : IsClosed (s : Set A)) : + elemental R x ≤ s ↔ x ∈ s := + ⟨fun h ↦ h (self_mem R x), fun h ↦ le_of_mem hs h⟩ + +instance isClosed (x : A) : IsClosed (elemental R x : Set A) := + isClosed_topologicalClosure _ + +instance [T2Space A] {x : A} : NonUnitalCommSemiring (elemental R x) := + nonUnitalCommSemiringTopologicalClosure _ + letI : NonUnitalCommSemiring (adjoin R {x}) := + NonUnitalAlgebra.adjoinNonUnitalCommSemiringOfComm R fun y hy z hz => by + rw [Set.mem_singleton_iff] at hy hz + rw [hy, hz] + fun _ _ => mul_comm _ _ + +instance {R A : Type*} [CommRing R] [NonUnitalRing A] + [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] + [TopologicalSpace A] [TopologicalRing A] [ContinuousConstSMul R A] + [T2Space A] {x : A} : NonUnitalCommRing (elemental R x) where + mul_comm := mul_comm + +instance {A : Type*} [UniformSpace A] [CompleteSpace A] [NonUnitalSemiring A] + [TopologicalSemiring A] [Module R A] [IsScalarTower R A A] + [SMulCommClass R A A] [ContinuousConstSMul R A] (x : A) : + CompleteSpace (elemental R x) := + isClosed_closure.completeSpace_coe + +/-- The coercion from an elemental algebra to the full algebra is a `IsClosedEmbedding`. -/ +theorem isClosedEmbedding_coe (x : A) : Topology.IsClosedEmbedding ((↑) : elemental R x → A) where + eq_induced := rfl + injective := Subtype.coe_injective + isClosed_range := by simpa using isClosed R x + +end elemental + +end NonUnitalAlgebra diff --git a/Mathlib/Topology/Algebra/NonUnitalStarAlgebra.lean b/Mathlib/Topology/Algebra/NonUnitalStarAlgebra.lean index 92452ae7c1b4a..5c74a83675c40 100644 --- a/Mathlib/Topology/Algebra/NonUnitalStarAlgebra.lean +++ b/Mathlib/Topology/Algebra/NonUnitalStarAlgebra.lean @@ -81,3 +81,71 @@ abbrev nonUnitalCommRingTopologicalClosure [T2Space A] (s : NonUnitalStarSubalge end Ring end NonUnitalStarSubalgebra + +namespace NonUnitalStarAlgebra + +open NonUnitalStarSubalgebra + +variable (R : Type*) {A : Type*} [CommSemiring R] [StarRing R] [NonUnitalSemiring A] [StarRing A] +variable [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] [StarModule R A] +variable [TopologicalSpace A] [TopologicalSemiring A] [ContinuousConstSMul R A] [ContinuousStar A] + +/-- The topological closure of the non-unital star subalgebra generated by a single element. -/ +def elemental (x : A) : NonUnitalStarSubalgebra R A := + adjoin R {x} |>.topologicalClosure + +namespace elemental + +@[aesop safe apply (rule_sets := [SetLike])] +theorem self_mem (x : A) : x ∈ elemental R x := + le_topologicalClosure _ <| self_mem_adjoin_singleton R x + +@[aesop safe apply (rule_sets := [SetLike])] +theorem star_self_mem (x : A) : star x ∈ elemental R x := + le_topologicalClosure _ <| star_self_mem_adjoin_singleton R x + +variable {R} in +theorem le_of_mem {x : A} {s : NonUnitalStarSubalgebra R A} (hs : IsClosed (s : Set A)) + (hx : x ∈ s) : elemental R x ≤ s := + topologicalClosure_minimal _ (adjoin_le <| by simpa using hx) hs + +variable {R} in +theorem le_iff_mem {x : A} {s : NonUnitalStarSubalgebra R A} (hs : IsClosed (s : Set A)) : + elemental R x ≤ s ↔ x ∈ s := + ⟨fun h ↦ h (self_mem R x), fun h ↦ le_of_mem hs h⟩ + +instance isClosed (x : A) : IsClosed (elemental R x : Set A) := + isClosed_topologicalClosure _ + +instance [T2Space A] {x : A} [IsStarNormal x] : NonUnitalCommSemiring (elemental R x) := + nonUnitalCommSemiringTopologicalClosure _ <| by + letI : NonUnitalCommSemiring (adjoin R {x}) := by + refine NonUnitalStarAlgebra.adjoinNonUnitalCommSemiringOfComm R ?_ ?_ + all_goals + intro y hy z hz + rw [Set.mem_singleton_iff] at hy hz + rw [hy, hz] + exact (star_comm_self' x).symm + exact fun _ _ => mul_comm _ _ + +instance {R A : Type*} [CommRing R] [StarRing R] [NonUnitalRing A] [StarRing A] + [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] [StarModule R A] + [TopologicalSpace A] [TopologicalRing A] [ContinuousConstSMul R A] [ContinuousStar A] + [T2Space A] {x : A} [IsStarNormal x] : NonUnitalCommRing (elemental R x) where + mul_comm := mul_comm + +instance {A : Type*} [UniformSpace A] [CompleteSpace A] [NonUnitalSemiring A] [StarRing A] + [TopologicalSemiring A] [ContinuousStar A] [Module R A] [IsScalarTower R A A] + [SMulCommClass R A A] [StarModule R A] [ContinuousConstSMul R A] (x : A) : + CompleteSpace (elemental R x) := + isClosed_closure.completeSpace_coe + +/-- The coercion from an elemental algebra to the full algebra is a `IsClosedEmbedding`. -/ +theorem isClosedEmbedding_coe (x : A) : Topology.IsClosedEmbedding ((↑) : elemental R x → A) where + eq_induced := rfl + injective := Subtype.coe_injective + isClosed_range := by simpa using isClosed R x + +end elemental + +end NonUnitalStarAlgebra From bdc1ed1f38009380f142b3500d72050c62faaf81 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Wed, 27 Nov 2024 03:15:31 +0000 Subject: [PATCH 336/829] =?UTF-8?q?feat:=20approximate=20units=20in=20C?= =?UTF-8?q?=E2=8B=86-algebras=20(#18506)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This constructs the canonical approximate unit in an arbitrary C⋆-algebra. This approximate unit is the filter corresponding to the basis of sections `{x | a ≤ x} ∩ closedBall 0 1`, where `0 ≤ a` and `‖a‖ < 1`. --- Mathlib.lean | 1 + .../CStarAlgebra/ApproximateUnit.lean | 221 +++++++++++++++++- .../SpecialFunctions/PosPart.lean | 55 +++++ .../ContinuousFunctionalCalculus/PosPart.lean | 26 +++ 4 files changed, 292 insertions(+), 11 deletions(-) create mode 100644 Mathlib/Analysis/CStarAlgebra/SpecialFunctions/PosPart.lean diff --git a/Mathlib.lean b/Mathlib.lean index 1fe86e9134e3b..0195cb1242e40 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1107,6 +1107,7 @@ import Mathlib.Analysis.CStarAlgebra.Module.Constructions import Mathlib.Analysis.CStarAlgebra.Module.Defs import Mathlib.Analysis.CStarAlgebra.Module.Synonym import Mathlib.Analysis.CStarAlgebra.Multiplier +import Mathlib.Analysis.CStarAlgebra.SpecialFunctions.PosPart import Mathlib.Analysis.CStarAlgebra.Spectrum import Mathlib.Analysis.CStarAlgebra.Unitization import Mathlib.Analysis.CStarAlgebra.lpSpace diff --git a/Mathlib/Analysis/CStarAlgebra/ApproximateUnit.lean b/Mathlib/Analysis/CStarAlgebra/ApproximateUnit.lean index 89f1b4ea09125..49852fd88a32f 100644 --- a/Mathlib/Analysis/CStarAlgebra/ApproximateUnit.lean +++ b/Mathlib/Analysis/CStarAlgebra/ApproximateUnit.lean @@ -5,15 +5,21 @@ Authors: Jireh Loreaux -/ import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Order import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Isometric +import Mathlib.Analysis.CStarAlgebra.SpecialFunctions.PosPart import Mathlib.Analysis.SpecialFunctions.ContinuousFunctionalCalculus.Rpow +import Mathlib.Topology.ApproximateUnit -/-! # Positive contractions in a C⋆-algebra form an approximate unit +/-! # Nonnegative contractions in a C⋆-algebra form an approximate unit -This file will show (although it does not yet) that the collection of positive contractions (of norm -strictly less than one) in a possibly non-unital C⋆-algebra form an approximate unit. The key step -is to show that this set is directed using the continuous functional calculus applied with the -functions `fun x : ℝ≥0, 1 - (1 + x)⁻¹` and `fun x : ℝ≥0, x * (1 - x)⁻¹`, which are inverses on -the interval `{x : ℝ≥0 | x < 1}`. +This file shows that the collection of positive contractions (of norm strictly less than one) +in a possibly non-unital C⋆-algebra form a directed set. The key step uses the continuous functional +calculus applied with the functions `fun x : ℝ≥0, 1 - (1 + x)⁻¹` and `fun x : ℝ≥0, x * (1 - x)⁻¹`, +which are inverses on the interval `{x : ℝ≥0 | x < 1}`. + +In addition, this file defines `IsIncreasingApproximateUnit` to be a filter `l` that is an +approximate unit contained in the closed unit ball of nonnegative elements. Every C⋆-algebra has +a filter generated by the sections `{x | a ≤ x} ∩ closedBall 0 1` for `0 ≤ a` and `‖a‖ < 1`, and +moreover, this filter is an increasing approximate unit. ## Main declarations @@ -22,11 +28,12 @@ the interval `{x : ℝ≥0 | x < 1}`. + `Set.InvOn.one_sub_one_add_inv`: the functions `f := fun x : ℝ≥0, 1 - (1 + x)⁻¹` and `g := fun x : ℝ≥0, x * (1 - x)⁻¹` are inverses on `{x : ℝ≥0 | x < 1}`. + `CStarAlgebra.directedOn_nonneg_ball`: the set `{x : A | 0 ≤ x} ∩ Metric.ball 0 1` is directed. - -## TODO - -+ Prove the approximate identity result by following Ken Davidson's proof in - "C*-Algebras by Example" ++ `Filter.IsIncreasingApproximateUnit`: a filter `l` is an *increasing approximate unit* if it is an + approximate unit contained in the closed unit ball of nonnegative elements. ++ `CStarAlgebra.approximateUnit`: the filter generated by the sections + `{x | a ≤ x} ∩ closedBall 0 1` for `0 ≤ a` with `‖a‖ < 1`. ++ `CStarAlgebra.increasingApproximateUnit`: the filter `CStarAlgebra.approximateUnit` is an + increasing approximate unit. -/ @@ -105,3 +112,195 @@ lemma CStarAlgebra.directedOn_nonneg_ball : have hab' : cfcₙ g a ≤ cfcₙ g a + cfcₙ g b := le_add_of_nonneg_right cfcₙ_nonneg_of_predicate exact CFC.monotoneOn_one_sub_one_add_inv cfcₙ_nonneg_of_predicate (cfcₙ_nonneg_of_predicate.trans hab') hab' + +section ApproximateUnit + +open Metric Filter Topology + +/-- An *increasing approximate unit* in a C⋆-algebra is an approximate unit contained in the +closed unit ball of nonnegative elements. -/ +structure Filter.IsIncreasingApproximateUnit (l : Filter A) extends l.IsApproximateUnit : Prop where + eventually_nonneg : ∀ᶠ x in l, 0 ≤ x + eventually_norm : ∀ᶠ x in l, ‖x‖ ≤ 1 + +namespace Filter.IsIncreasingApproximateUnit + +omit [StarOrderedRing A] in +lemma eventually_nnnorm {l : Filter A} (hl : l.IsIncreasingApproximateUnit) : + ∀ᶠ x in l, ‖x‖₊ ≤ 1 := + hl.eventually_norm + +lemma eventually_isSelfAdjoint {l : Filter A} (hl : l.IsIncreasingApproximateUnit) : + ∀ᶠ x in l, IsSelfAdjoint x := + hl.eventually_nonneg.mp <| .of_forall fun _ ↦ IsSelfAdjoint.of_nonneg + +lemma eventually_star_eq {l : Filter A} (hl : l.IsIncreasingApproximateUnit) : + ∀ᶠ x in l, star x = x := + hl.eventually_isSelfAdjoint.mp <| .of_forall fun _ ↦ IsSelfAdjoint.star_eq + +end Filter.IsIncreasingApproximateUnit + +namespace CStarAlgebra + +open Submodule in +/-- To show that `l` is a one-sided approximate unit for `A`, it suffices to verify it only for +`m : A` with `0 ≤ m` and `‖m‖ < 1`. -/ +lemma tendsto_mul_right_of_forall_nonneg_tendsto {l : Filter A} + (h : ∀ m, 0 ≤ m → ‖m‖ < 1 → Tendsto (· * m) l (𝓝 m)) (m : A) : + Tendsto (· * m) l (𝓝 m) := by + obtain ⟨n, c, x, rfl⟩ := mem_span_set'.mp <| by + show m ∈ span ℂ ({x | 0 ≤ x} ∩ ball 0 1) + simp [span_nonneg_inter_unitBall] + simp_rw [Finset.mul_sum] + refine tendsto_finset_sum _ fun i _ ↦ ?_ + simp_rw [mul_smul_comm] + exact tendsto_const_nhds.smul <| h (x i) (x i).2.1 <| by simpa using (x i).2.2 + +omit [PartialOrder A] in +/-- Multiplication on the left by `m` tends to `𝓝 m` if and only if multiplication on the right +does, provided the elements are eventually selfadjoint along the filter `l`. -/ +lemma tendsto_mul_left_iff_tendsto_mul_right {l : Filter A} (hl : ∀ᶠ x in l, IsSelfAdjoint x) : + (∀ m, Tendsto (m * ·) l (𝓝 m)) ↔ (∀ m, Tendsto (· * m) l (𝓝 m)) := by + refine ⟨fun h m ↦ ?_, fun h m ↦ ?_⟩ + all_goals + apply (star_star m ▸ (continuous_star.tendsto _ |>.comp <| h (star m))).congr' + filter_upwards [hl] with x hx + simp [hx.star_eq] + +variable (A) + +/-- The sections of positive strict contractions form a filter basis. -/ +lemma isBasis_nonneg_sections : + IsBasis (fun x : A ↦ 0 ≤ x ∧ ‖x‖ < 1) ({x | · ≤ x}) where + nonempty := ⟨0, by simp⟩ + inter {x y} hx hy := by + peel directedOn_nonneg_ball x (by simpa) y (by simpa) with z hz + exact ⟨by simpa using hz.1, fun a ha ↦ ⟨hz.2.1.trans ha, hz.2.2.trans ha⟩⟩ + +/-- The canonical approximate unit in a C⋆-algebra generated by the basis of sets +`{x | a ≤ x} ∩ closedBall 0 1` for `0 ≤ a`. See also `CStarAlgebra.hasBasis_approximateUnit`. -/ +def approximateUnit : Filter A := + (isBasis_nonneg_sections A).filter ⊓ 𝓟 (closedBall 0 1) + +/-- The canonical approximate unit in a C⋆-algebra has a basis of sets +`{x | a ≤ x} ∩ closedBall 0 1` for `0 ≤ a`. -/ +lemma hasBasis_approximateUnit : + (approximateUnit A).HasBasis (fun x : A ↦ 0 ≤ x ∧ ‖x‖ < 1) ({x | · ≤ x} ∩ closedBall 0 1) := + isBasis_nonneg_sections A |>.hasBasis.inf_principal (closedBall 0 1) + +/-- This is a common reasoning sequence in C⋆-algebra theory. If `0 ≤ x ≤ y ≤ 1`, then the norm +of `z - y * z` is controled by the norm of `star z * (1 - x) * z`, which is advantageous because the +latter is nonnegative. This is a key step in establishing the existence of an increasing approximate +unit in general C⋆-algebras. -/ +lemma nnnorm_sub_mul_self_le {A : Type*} [CStarAlgebra A] [PartialOrder A] [StarOrderedRing A] + {x y : A} (z : A) (hx₀ : 0 ≤ x) (hy : y ∈ Set.Icc x 1) {c : ℝ≥0} + (h : ‖star z * (1 - x) * z‖₊ ≤ c ^ 2) : + ‖z - y * z‖₊ ≤ c := by + nth_rw 1 [← one_mul z] + rw [← sqrt_sq c, le_sqrt_iff_sq_le, ← sub_mul, sq, ← CStarRing.nnnorm_star_mul_self] + simp only [star_mul, star_sub, star_one] + have hy₀ : y ∈ Set.Icc 0 1 := ⟨hx₀.trans hy.1, hy.2⟩ + have hy' : 1 - y ∈ Set.Icc 0 1 := Set.sub_mem_Icc_zero_iff_right.mpr hy₀ + rw [hy₀.1.star_eq, ← mul_assoc, mul_assoc (star _), ← sq] + refine nnnorm_le_nnnorm_of_nonneg_of_le (conjugate_nonneg (pow_nonneg hy'.1 2) _) ?_ |>.trans h + refine conjugate_le_conjugate ?_ _ + trans (1 - y) + · simpa using pow_antitone hy'.1 hy'.2 one_le_two + · gcongr + exact hy.1 + +/-- A variant of `nnnorm_sub_mul_self_le` which uses `‖·‖` instead of `‖·‖₊`. -/ +lemma norm_sub_mul_self_le {A : Type*} [CStarAlgebra A] [PartialOrder A] [StarOrderedRing A] + {x y : A} (z : A) (hx₀ : 0 ≤ x) (hy : y ∈ Set.Icc x 1) + {c : ℝ} (hc : 0 ≤ c) (h : ‖star z * (1 - x) * z‖ ≤ c ^ 2) : + ‖z - y * z‖ ≤ c := + nnnorm_sub_mul_self_le z hx₀ hy h (c := ⟨c, hc⟩) + +variable {A} in +/-- A variant of `norm_sub_mul_self_le` for non-unital algebras that passes to the unitization. -/ +lemma norm_sub_mul_self_le_of_inr {x y : A} (z : A) (hx₀ : 0 ≤ x) (hxy : x ≤ y) (hy₁ : ‖y‖ ≤ 1) + {c : ℝ} (hc : 0 ≤ c) (h : ‖star (z : A⁺¹) * (1 - x) * z‖ ≤ c ^ 2) : + ‖z - y * z‖ ≤ c := by + rw [← norm_inr (𝕜 := ℂ), inr_sub, inr_mul] + refine norm_sub_mul_self_le _ ?_ ?_ hc h + · rwa [inr_nonneg_iff] + · have hy := hx₀.trans hxy + rw [Set.mem_Icc, inr_le_iff _ _ hx₀.isSelfAdjoint hy.isSelfAdjoint, + ← norm_le_one_iff_of_nonneg _, norm_inr] + exact ⟨hxy, hy₁⟩ + +variable {A} in +/-- This shows `CStarAlgebra.approximateUnit` is a one-sided approximate unit, but this is marked +`private` because it is only used to prove `CStarAlgebra.increasingApproximateUnit`. -/ +private lemma tendsto_mul_right_approximateUnit (m : A) : + Tendsto (· * m) (approximateUnit A) (𝓝 m) := by + refine tendsto_mul_right_of_forall_nonneg_tendsto (fun m hm₁ hm₂ ↦ ?_) m + rw [(hasBasis_approximateUnit A).tendsto_iff nhds_basis_closedBall] + intro ε hε + lift ε to ℝ≥0 using hε.le + rw [coe_pos] at hε + refine ⟨cfcₙ (fun y : ℝ≥0 ↦ 1 - (1 + y)⁻¹) (ε⁻¹ ^ 2 • m), + ⟨cfcₙ_nonneg_of_predicate, norm_cfcₙ_one_sub_one_add_inv_lt_one (ε⁻¹ ^ 2 • m)⟩, ?_⟩ + rintro x ⟨(hx₁ : _ ≤ x), hx₂⟩ + simp only [mem_closedBall, dist_eq_norm', zero_sub, norm_neg] at hx₂ ⊢ + rw [← coe_nnnorm, coe_le_coe] + have hx₀ : 0 ≤ x := cfcₙ_nonneg_of_predicate.trans hx₁ + rw [← inr_le_iff _ _ (.of_nonneg cfcₙ_nonneg_of_predicate) (.of_nonneg hx₀), + nnreal_cfcₙ_eq_cfc_inr _ _ (by simp [tsub_self]), inr_smul] at hx₁ + rw [← norm_inr (𝕜 := ℂ)] at hm₂ hx₂ + rw [← inr_nonneg_iff] at hx₀ hm₁ + rw [← nnnorm_inr (𝕜 := ℂ), inr_sub, inr_mul] + generalize (x : A⁺¹) = x, (m : A⁺¹) = m at * + set g : ℝ≥0 → ℝ≥0 := fun y ↦ 1 - (1 + y)⁻¹ + have hg : Continuous g := by + rw [continuous_iff_continuousOn_univ] + fun_prop (disch := intro _ _; positivity) + have hg' : ContinuousOn (fun y ↦ (1 + ε⁻¹ ^ 2 • y)⁻¹) (spectrum ℝ≥0 m) := + ContinuousOn.inv₀ (by fun_prop) fun _ _ ↦ by positivity + have hx : x ∈ Set.Icc 0 1 := mem_Icc_iff_norm_le_one.mpr ⟨hx₀, hx₂⟩ + have hx' : x ∈ Set.Icc _ 1 := ⟨hx₁, hx.2⟩ + refine nnnorm_sub_mul_self_le m cfc_nonneg_of_predicate hx' ?_ + suffices star m * (1 - cfc g (ε⁻¹ ^ 2 • m)) * m = + cfc (fun y : ℝ≥0 ↦ y * (1 + ε⁻¹ ^ 2 • y)⁻¹ * y) m by + rw [this] + refine nnnorm_cfc_nnreal_le fun y hy ↦ ?_ + field_simp + calc + y * ε ^ 2 * y / (ε ^ 2 + y) ≤ ε ^ 2 * 1 := by + rw [mul_div_assoc] + gcongr + · refine mul_le_of_le_one_left (zero_le _) ?_ + have hm' := hm₂.le + rw [norm_le_one_iff_of_nonneg m hm₁, ← cfc_id' ℝ≥0 m, ← cfc_one (R := ℝ≥0) m, + cfc_nnreal_le_iff _ _ _ (QuasispectrumRestricts.nnreal_of_nonneg hm₁)] at hm' + exact hm' y hy + · exact div_le_one (by positivity) |>.mpr le_add_self + _ = ε ^ 2 := mul_one _ + rw [cfc_mul _ _ m (continuousOn_id' _ |>.mul hg') (continuousOn_id' _), + cfc_mul _ _ m (continuousOn_id' _) hg', cfc_id' .., hm₁.star_eq] + congr + rw [← cfc_one (R := ℝ≥0) m, ← cfc_comp_smul _ _ _ hg.continuousOn hm₁, + ← cfc_tsub _ _ m (by simp [g]) hm₁ (by fun_prop) (Continuous.continuousOn <| by fun_prop)] + refine cfc_congr (fun y _ ↦ ?_) + simp [g, tsub_tsub_cancel_of_le] + +/-- The filter `CStarAlgebra.approximateUnit` generated by the sections +`{x | a ≤ x} ∩ closedBall 0 1` for `0 ≤ a` forms an increasing approximate unit. -/ +lemma increasingApproximateUnit : + IsIncreasingApproximateUnit (approximateUnit A) where + tendsto_mul_left := by + rw [tendsto_mul_left_iff_tendsto_mul_right] + · exact tendsto_mul_right_approximateUnit + · rw [(hasBasis_approximateUnit A).eventually_iff] + peel (hasBasis_approximateUnit A).ex_mem with x hx + exact ⟨hx, fun y hy ↦ (hx.1.trans hy.1).isSelfAdjoint⟩ + tendsto_mul_right := tendsto_mul_right_approximateUnit + eventually_nonneg := .filter_mono inf_le_left <| + (isBasis_nonneg_sections A).hasBasis.eventually_iff.mpr ⟨0, by simp⟩ + eventually_norm := .filter_mono inf_le_right <| by simp + neBot := hasBasis_approximateUnit A |>.neBot_iff.mpr + fun hx ↦ ⟨_, ⟨le_rfl, by simpa using hx.2.le⟩⟩ + +end CStarAlgebra + +end ApproximateUnit diff --git a/Mathlib/Analysis/CStarAlgebra/SpecialFunctions/PosPart.lean b/Mathlib/Analysis/CStarAlgebra/SpecialFunctions/PosPart.lean new file mode 100644 index 0000000000000..de68421d1db20 --- /dev/null +++ b/Mathlib/Analysis/CStarAlgebra/SpecialFunctions/PosPart.lean @@ -0,0 +1,55 @@ +/- +Copyright (c) 2024 Jireh Loreaux. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jireh Loreaux +-/ +import Mathlib.Analysis.SpecialFunctions.ContinuousFunctionalCalculus.PosPart +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Instances + +/-! # C⋆-algebraic facts about `a⁺` and `a⁻`. -/ + +variable {A : Type*} [NonUnitalCStarAlgebra A] [PartialOrder A] [StarOrderedRing A] + +namespace CStarAlgebra + +section SpanNonneg + +open Submodule + +/-- A C⋆-algebra is spanned by nonnegative elements of norm at most `r` -/ +lemma span_nonneg_inter_closedBall {r : ℝ} (hr : 0 < r) : + span ℂ ({x : A | 0 ≤ x} ∩ Metric.closedBall 0 r) = ⊤ := by + rw [eq_top_iff, ← span_nonneg, span_le] + intro x hx + obtain (rfl | hx_pos) := eq_zero_or_norm_pos x + · exact zero_mem _ + · suffices (r * ‖x‖⁻¹ : ℂ)⁻¹ • ((r * ‖x‖⁻¹ : ℂ) • x) = x by + rw [← this] + refine smul_mem _ _ (subset_span <| Set.mem_inter ?_ ?_) + · norm_cast + exact smul_nonneg (by positivity) hx + · simp [mul_smul, norm_smul, abs_of_pos hr, inv_mul_cancel₀ hx_pos.ne'] + apply inv_smul_smul₀ + norm_cast + positivity + +/-- A C⋆-algebra is spanned by nonnegative elements of norm less than `r`. -/ +lemma span_nonneg_inter_ball {r : ℝ} (hr : 0 < r) : + span ℂ ({x : A | 0 ≤ x} ∩ Metric.ball 0 r) = ⊤ := by + rw [eq_top_iff, ← span_nonneg_inter_closedBall (half_pos hr)] + gcongr + exact Metric.closedBall_subset_ball <| half_lt_self hr + +/-- A C⋆-algebra is spanned by nonnegative contractions. -/ +lemma span_nonneg_inter_unitClosedBall : + span ℂ ({x : A | 0 ≤ x} ∩ Metric.closedBall 0 1) = ⊤ := + span_nonneg_inter_closedBall zero_lt_one + +/-- A C⋆-algebra is spanned by nonnegative strict contractions. -/ +lemma span_nonneg_inter_unitBall : + span ℂ ({x : A | 0 ≤ x} ∩ Metric.ball 0 1) = ⊤ := + span_nonneg_inter_ball zero_lt_one + +end SpanNonneg + +end CStarAlgebra diff --git a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/PosPart.lean b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/PosPart.lean index 396d6e6043dca..959183c53b686 100644 --- a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/PosPart.lean +++ b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/PosPart.lean @@ -226,3 +226,29 @@ lemma posPart_negPart_unique {a b c : A} (habc : a = b - c) (hbc : b * c = 0) simp [Subtype.val_injective.extend_apply, f] end CFC + +section SpanNonneg + +variable {A : Type*} [NonUnitalRing A] [Module ℂ A] [SMulCommClass ℂ A A] [IsScalarTower ℂ A A] +variable [StarRing A] [TopologicalSpace A] [StarModule ℂ A] +variable [NonUnitalContinuousFunctionalCalculus ℝ (IsSelfAdjoint : A → Prop)] +variable [PartialOrder A] [StarOrderedRing A] + +open Submodule Complex +open scoped ComplexStarModule + +lemma CStarAlgebra.linear_combination_nonneg (x : A) : + ((ℜ x : A)⁺ - (ℜ x : A)⁻) + (I • (ℑ x : A)⁺ - I • (ℑ x : A)⁻) = x := by + rw [CFC.posPart_sub_negPart _ (ℜ x).2, ← smul_sub, CFC.posPart_sub_negPart _ (ℑ x).2, + realPart_add_I_smul_imaginaryPart x] + +/-- A C⋆-algebra is spanned by its nonnegative elements. -/ +lemma CStarAlgebra.span_nonneg : Submodule.span ℂ {a : A | 0 ≤ a} = ⊤ := by + refine eq_top_iff.mpr fun x _ => ?_ + rw [← CStarAlgebra.linear_combination_nonneg x] + apply_rules [sub_mem, Submodule.smul_mem, add_mem] + all_goals + refine subset_span ?_ + first | apply CFC.negPart_nonneg | apply CFC.posPart_nonneg + +end SpanNonneg From 8a7a0b7d493c44f352e65718307fce1c052d20e2 Mon Sep 17 00:00:00 2001 From: judithludwig Date: Wed, 27 Nov 2024 08:00:36 +0000 Subject: [PATCH 337/829] refactor(LinearAlgebra/Matrix/GeneralLinearGroup): Redefine `Matrix.SpecialLinearGroup.toGL`and remove duplicate `toLinear` (#19523) This PR redefines `Matrix.SpecialLinearGroup.toGL`, as the multiplicative map from the matrix group $$\mathrm{SL}_n(R)$$ to the matrix group $$\mathrm{GL}_n(R)$$, following a [Zulip discussion](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/Matrix.2ESpecialLinearGroup.2EtoGL). It also removes a duplicate of `Matrix.GeneralLinearGroup.toLin`. Co-authored-by: judithludwig <150042763+judithludwig@users.noreply.github.com> --- .../Complex/UpperHalfPlane/Basic.lean | 6 ++-- .../Matrix/GeneralLinearGroup/Defs.lean | 31 +++++++++---------- .../Matrix/SpecialLinearGroup.lean | 9 ------ .../LinearAlgebra/Matrix/ToLinearEquiv.lean | 2 +- Mathlib/NumberTheory/Modular.lean | 6 ++-- 5 files changed, 22 insertions(+), 32 deletions(-) diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean index 6a3125c87d0c9..3a538cfd9d8ac 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean @@ -364,12 +364,12 @@ end RealAddAction /- these next few lemmas are *not* flagged `@simp` because of the constructors on the RHS; instead we use the versions with coercions to `ℂ` as simp lemmas instead. -/ theorem modular_S_smul (z : ℍ) : ModularGroup.S • z = mk (-z : ℂ)⁻¹ z.im_inv_neg_coe_pos := by - rw [specialLinearGroup_apply]; simp [ModularGroup.S, neg_div, inv_neg, coeToGL] + rw [specialLinearGroup_apply]; simp [ModularGroup.S, neg_div, inv_neg, toGL] theorem modular_T_zpow_smul (z : ℍ) (n : ℤ) : ModularGroup.T ^ n • z = (n : ℝ) +ᵥ z := by rw [UpperHalfPlane.ext_iff, coe_vadd, add_comm, specialLinearGroup_apply, coe_mk] -- Porting note: added `coeToGL` and merged `rw` and `simp` - simp [coeToGL, ModularGroup.coe_T_zpow, + simp [toGL, ModularGroup.coe_T_zpow, of_apply, cons_val_zero, algebraMap.coe_one, Complex.ofReal_one, one_mul, cons_val_one, head_cons, algebraMap.coe_zero, zero_mul, zero_add, div_one] @@ -383,7 +383,7 @@ theorem exists_SL2_smul_eq_of_apply_zero_one_eq_zero (g : SL(2, ℝ)) (hc : g 1 ext1 ⟨z, hz⟩; ext1 suffices ↑a * z * a + b * a = b * a + a * a * z by -- Porting note: added `coeToGL` and merged `rw` and `simpa` - simpa [coeToGL, specialLinearGroup_apply, add_mul] + simpa [toGL, specialLinearGroup_apply, add_mul] ring theorem exists_SL2_smul_eq_of_apply_zero_one_ne_zero (g : SL(2, ℝ)) (hc : g 1 0 ≠ 0) : diff --git a/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean b/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean index 62ce0b6d7abc0..6a59c9e32abf2 100644 --- a/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean +++ b/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean @@ -5,6 +5,7 @@ Authors: Chris Birkbeck -/ import Mathlib.LinearAlgebra.Matrix.NonsingularInverse import Mathlib.LinearAlgebra.Matrix.SpecialLinearGroup +import Mathlib.LinearAlgebra.GeneralLinearGroup import Mathlib.Algebra.Ring.Subring.Units /-! @@ -69,7 +70,8 @@ def det : GL n R →* Rˣ where map_one' := Units.ext det_one map_mul' _ _ := Units.ext <| det_mul _ _ -/-- The `GL n R` and `Matrix.GeneralLinearGroup R n` groups are multiplicatively equivalent -/ +/-- The groups `GL n R` (notation for `Matrix.GeneralLinearGroup n R`) and +`LinearMap.GeneralLinearGroup R (n → R)` are multiplicatively equivalent -/ def toLin : GL n R ≃* LinearMap.GeneralLinearGroup R (n → R) := Units.mapEquiv toLinAlgEquiv'.toMulEquiv @@ -108,22 +110,19 @@ theorem coe_inv : ↑A⁻¹ = (↑A : Matrix n n R)⁻¹ := letI := A.invertible invOf_eq_nonsing_inv (↑A : Matrix n n R) -/-- An element of the matrix general linear group on `(n) [Fintype n]` can be considered as an -element of the endomorphism general linear group on `n → R`. -/ -def toLinear : GeneralLinearGroup n R ≃* LinearMap.GeneralLinearGroup R (n → R) := - Units.mapEquiv Matrix.toLinAlgEquiv'.toRingEquiv.toMulEquiv +@[deprecated (since := "2024-11-26")] alias toLinear := toLin -- Note that without the `@` and `‹_›`, Lean infers `fun a b ↦ _inst a b` instead of `_inst` as the -- decidability argument, which prevents `simp` from obtaining the instance by unification. -- These `fun a b ↦ _inst a b` terms also appear in the type of `A`, but simp doesn't get confused -- by them so for now we do not care. @[simp] -theorem coe_toLinear : (@toLinear n ‹_› ‹_› _ _ A : (n → R) →ₗ[R] n → R) = Matrix.mulVecLin A := +theorem coe_toLin : (@toLin n ‹_› ‹_› _ _ A : (n → R) →ₗ[R] n → R) = Matrix.mulVecLin A := rfl -- Porting note: is inserting toLinearEquiv here correct? @[simp] -theorem toLinear_apply (v : n → R) : (toLinear A).toLinearEquiv v = Matrix.mulVecLin (↑A) v := +theorem toLin_apply (v : n → R) : (toLin A).toLinearEquiv v = Matrix.mulVecLin (↑A) v := rfl end CoeLemmas @@ -195,17 +194,17 @@ namespace SpecialLinearGroup variable {n : Type u} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] --- Porting note: added implementation for the Coe -/-- The map from SL(n) to GL(n) underlying the coercion, forgetting the value of the determinant. --/ -@[coe] -def coeToGL (A : SpecialLinearGroup n R) : GL n R := - ⟨↑A, ↑A⁻¹, - congr_arg ((↑) : _ → Matrix n n R) (mul_inv_cancel A), - congr_arg ((↑) : _ → Matrix n n R) (inv_mul_cancel A)⟩ + +/-- `toGL` is the map from the special linear group to the general linear group. -/ +def toGL : Matrix.SpecialLinearGroup n R →* Matrix.GeneralLinearGroup n R where + toFun A := ⟨↑A, ↑A⁻¹, congr_arg (·.1) (mul_inv_cancel A), congr_arg (·.1) (inv_mul_cancel A)⟩ + map_one' := Units.ext rfl + map_mul' _ _ := Units.ext rfl + +@[deprecated (since := "2024-11-26")] alias coeToGL := toGL instance hasCoeToGeneralLinearGroup : Coe (SpecialLinearGroup n R) (GL n R) := - ⟨coeToGL⟩ + ⟨toGL⟩ @[simp] theorem coeToGL_det (g : SpecialLinearGroup n R) : diff --git a/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean b/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean index be6c0ddf6440a..59bff4a8773bd 100644 --- a/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean +++ b/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean @@ -3,7 +3,6 @@ Copyright (c) 2020 Anne Baanen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen, Wen Yang -/ -import Mathlib.LinearAlgebra.GeneralLinearGroup import Mathlib.LinearAlgebra.Matrix.Adjugate import Mathlib.LinearAlgebra.Matrix.Transvection import Mathlib.RingTheory.RootsOfUnity.Basic @@ -211,14 +210,6 @@ theorem toLin'_injective : Function.Injective ↑(toLin' : SpecialLinearGroup n R →* (n → R) ≃ₗ[R] n → R) := fun _ _ h => Subtype.coe_injective <| Matrix.toLin'.injective <| LinearEquiv.toLinearMap_injective.eq_iff.mpr h -/-- `toGL` is the map from the special linear group to the general linear group -/ -def toGL : SpecialLinearGroup n R →* GeneralLinearGroup R (n → R) := - (GeneralLinearGroup.generalLinearEquiv _ _).symm.toMonoidHom.comp toLin' - --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation -theorem coe_toGL (A : SpecialLinearGroup n R) : SpecialLinearGroup.toGL A = A.toLin'.toLinearMap := - rfl - variable {S : Type*} [CommRing S] /-- A ring homomorphism from `R` to `S` induces a group homomorphism from diff --git a/Mathlib/LinearAlgebra/Matrix/ToLinearEquiv.lean b/Mathlib/LinearAlgebra/Matrix/ToLinearEquiv.lean index fefbfcaeaffce..c8f11ba565e7a 100644 --- a/Mathlib/LinearAlgebra/Matrix/ToLinearEquiv.lean +++ b/Mathlib/LinearAlgebra/Matrix/ToLinearEquiv.lean @@ -50,7 +50,7 @@ See `Matrix.toLinearEquiv` for the same map on arbitrary modules. -/ def toLinearEquiv' (P : Matrix n n R) (_ : Invertible P) : (n → R) ≃ₗ[R] n → R := GeneralLinearGroup.generalLinearEquiv _ _ <| - Matrix.GeneralLinearGroup.toLinear <| unitOfInvertible P + Matrix.GeneralLinearGroup.toLin <| unitOfInvertible P @[simp] theorem toLinearEquiv'_apply (P : Matrix n n R) (h : Invertible P) : diff --git a/Mathlib/NumberTheory/Modular.lean b/Mathlib/NumberTheory/Modular.lean index 59b083317b35b..acad10c21c22c 100644 --- a/Mathlib/NumberTheory/Modular.lean +++ b/Mathlib/NumberTheory/Modular.lean @@ -175,7 +175,7 @@ def lcRow0Extend {cd : Fin 2 → ℤ} (hcd : IsCoprime (cd 0) (cd 1)) : ![by refine LinearMap.GeneralLinearGroup.generalLinearEquiv ℝ (Fin 2 → ℝ) - (GeneralLinearGroup.toLinear (planeConformalMatrix (cd 0 : ℝ) (-(cd 1 : ℝ)) ?_)) + (GeneralLinearGroup.toLin (planeConformalMatrix (cd 0 : ℝ) (-(cd 1 : ℝ)) ?_)) norm_cast rw [neg_sq] exact hcd.sq_add_sq_ne_zero, LinearEquiv.refl ℝ (Fin 2 → ℝ)] @@ -210,12 +210,12 @@ theorem tendsto_lcRow0 {cd : Fin 2 → ℤ} (hcd : IsCoprime (cd 0) (cd 1)) : · simp only [Fin.isValue, Int.cast_one, map_apply_coe, RingHom.mapMatrix_apply, Int.coe_castRingHom, lcRow0_apply, map_apply, Fin.zero_eta, id_eq, Function.comp_apply, of_apply, cons_val', cons_val_zero, empty_val', cons_val_fin_one, lcRow0Extend_apply, - LinearMap.GeneralLinearGroup.coeFn_generalLinearEquiv, GeneralLinearGroup.coe_toLinear, + LinearMap.GeneralLinearGroup.coeFn_generalLinearEquiv, GeneralLinearGroup.coe_toLin, val_planeConformalMatrix, neg_neg, mulVecLin_apply, mulVec, dotProduct, Fin.sum_univ_two, cons_val_one, head_cons, mB, f₁] · convert congr_arg (fun n : ℤ => (-n : ℝ)) g.det_coe.symm using 1 simp only [Fin.zero_eta, id_eq, Function.comp_apply, lcRow0Extend_apply, cons_val_zero, - LinearMap.GeneralLinearGroup.coeFn_generalLinearEquiv, GeneralLinearGroup.coe_toLinear, + LinearMap.GeneralLinearGroup.coeFn_generalLinearEquiv, GeneralLinearGroup.coe_toLin, mulVecLin_apply, mulVec, dotProduct, det_fin_two, f₁] simp only [Fin.isValue, Fin.mk_one, val_planeConformalMatrix, neg_neg, of_apply, cons_val', empty_val', cons_val_fin_one, cons_val_one, head_fin_const, map_apply, Fin.sum_univ_two, From 7e86316a6a2b02c4957950b7ee550b17506485f5 Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Wed, 27 Nov 2024 08:55:06 +0000 Subject: [PATCH 338/829] feat(FieldTheory/LinearDisjoint): definition and basic properties of linearly disjoint of fields (#9651) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds basics about the linearly disjoint fields. Main definitions: - `IntermediateField.LinearDisjoint`: an intermediate field `A` of `E / F` and an abstract field `L` between `E / F` are linearly disjoint over `F`, if they are linearly disjoint as subalgebras (`Subalgebra.LinearDisjoint`). Main results: Equivalent characterization of linearly disjointness: - `IntermediateField.LinearDisjoint.linearIndependent_left`: if `A` and `L` are linearly disjoint, then any `F`-linearly independent family on `A` remains linearly independent over `L`. - `IntermediateField.LinearDisjoint.of_basis_left`: conversely, if there exists an `F`-basis of `A` which remains linearly independent over `L`, then `A` and `L` are linearly disjoint. - `IntermediateField.LinearDisjoint.linearIndependent_right`: if `A` and `L` are linearly disjoint, then any `F`-linearly independent family on `L` remains linearly independent over `A`. - `IntermediateField.LinearDisjoint.of_basis_right`: conversely, if there exists an `F`-basis of `L` which remains linearly independent over `A`, then `A` and `L` are linearly disjoint. - `IntermediateField.LinearDisjoint.linearIndependent_mul`: if `A` and `L` are linearly disjoint, then for any family of `F`-linearly independent elements `{ a_i }` of `A`, and any family of `F`-linearly independent elements `{ b_j }` of `L`, the family `{ a_i * b_j }` in `S` is also `F`-linearly independent. - `IntermediateField.LinearDisjoint.of_basis_mul`: conversely, if `{ a_i }` is an `F`-basis of `A`, if `{ b_j }` is an `F`-basis of `L`, such that the family `{ a_i * b_j }` in `E` is `F`-linearly independent, then `A` and `L` are linearly disjoint. Other main results: - `IntermediateField.LinearDisjoint.symm`, `IntermediateField.linearDisjoint_symm`: linearly disjoint is symmetric. - `IntermediateField.LinearDisjoint.rank_sup_of_isAlgebraic`, `IntermediateField.LinearDisjoint.finrank_sup`: if `A` and `B` are linearly disjoint, then the rank of `A ⊔ B` is equal to the product of the rank of `A` and `B`. **TODO:** remove the algebraic assumptions (the proof becomes complicated). - `IntermediateField.LinearDisjoint.of_finrank_sup`: conversely, if `A` and `B` are finite extensions, such that rank of `A ⊔ B` is equal to the product of the rank of `A` and `B`, then `A` and `B` are linearly disjoint. - `IntermediateField.LinearDisjoint.of_finrank_coprime`: if the rank of `A` and `B` are coprime, then `A` and `B` are linearly disjoint. - `IntermediateField.LinearDisjoint.inf_eq_bot`: if `A` and `B` are linearly disjoint, then they are disjoint. --- Mathlib.lean | 1 + Mathlib/FieldTheory/LinearDisjoint.lean | 310 ++++++++++++++++++++++ Mathlib/LinearAlgebra/LinearDisjoint.lean | 26 +- Mathlib/RingTheory/LinearDisjoint.lean | 20 +- 4 files changed, 334 insertions(+), 23 deletions(-) create mode 100644 Mathlib/FieldTheory/LinearDisjoint.lean diff --git a/Mathlib.lean b/Mathlib.lean index 0195cb1242e40..26f7426ffef46 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2971,6 +2971,7 @@ import Mathlib.FieldTheory.JacobsonNoether import Mathlib.FieldTheory.KrullTopology import Mathlib.FieldTheory.KummerExtension import Mathlib.FieldTheory.Laurent +import Mathlib.FieldTheory.LinearDisjoint import Mathlib.FieldTheory.Minpoly.Basic import Mathlib.FieldTheory.Minpoly.Field import Mathlib.FieldTheory.Minpoly.IsConjRoot diff --git a/Mathlib/FieldTheory/LinearDisjoint.lean b/Mathlib/FieldTheory/LinearDisjoint.lean new file mode 100644 index 0000000000000..3ad86d6638858 --- /dev/null +++ b/Mathlib/FieldTheory/LinearDisjoint.lean @@ -0,0 +1,310 @@ +/- +Copyright (c) 2024 Jz Pan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jz Pan +-/ +import Mathlib.FieldTheory.Adjoin +import Mathlib.RingTheory.LinearDisjoint + +/-! + +# Linearly disjoint fields + +This file contains basics about the linearly disjoint fields. + +## Linear disjoint intermediate fields + +We adapt the definitions in . +See the file `Mathlib/LinearAlgebra/LinearDisjoint.lean` for details. + +### Main definitions + +- `IntermediateField.LinearDisjoint`: an intermediate field `A` of `E / F` + and an abstract field `L` between `E / F` + (as a special case, two intermediate fields) are linearly disjoint over `F`, + if they are linearly disjoint as subalgebras (`Subalgebra.LinearDisjoint`). + +### Implementation notes + +The `Subalgebra.LinearDisjoint` is stated for two `Subalgebra`s. The original design of +`IntermediateField.LinearDisjoint` is also stated for two `IntermediateField`s +(see `IntermediateField.linearDisjoint_iff'` for the original statement). +But it's probably useful if one of them can be generalized to an abstract field +(see ). +This leads to the current design of `IntermediateField.LinearDisjoint` +which is for one `IntermediateField` and one abstract field. +It is not generalized to two abstract fields as this will break the dot notation. + +### Main results + +Equivalent characterization of linear disjointness: + +- `IntermediateField.LinearDisjoint.linearIndependent_left`: + if `A` and `L` are linearly disjoint, then any `F`-linearly independent family on `A` remains + linearly independent over `L`. + +- `IntermediateField.LinearDisjoint.of_basis_left`: + conversely, if there exists an `F`-basis of `A` which remains linearly independent over `L`, then + `A` and `L` are linearly disjoint. + +- `IntermediateField.LinearDisjoint.linearIndependent_right`: + `IntermediateField.LinearDisjoint.linearIndependent_right'`: + if `A` and `L` are linearly disjoint, then any `F`-linearly independent family on `L` remains + linearly independent over `A`. + +- `IntermediateField.LinearDisjoint.of_basis_right`: + `IntermediateField.LinearDisjoint.of_basis_right'`: + conversely, if there exists an `F`-basis of `L` which remains linearly independent over `A`, then + `A` and `L` are linearly disjoint. + +- `IntermediateField.LinearDisjoint.linearIndependent_mul`: + `IntermediateField.LinearDisjoint.linearIndependent_mul'`: + if `A` and `L` are linearly disjoint, then for any family of + `F`-linearly independent elements `{ a_i }` of `A`, and any family of + `F`-linearly independent elements `{ b_j }` of `L`, the family `{ a_i * b_j }` in `S` is + also `F`-linearly independent. + +- `IntermediateField.LinearDisjoint.of_basis_mul`: + `IntermediateField.LinearDisjoint.of_basis_mul'`: + conversely, if `{ a_i }` is an `F`-basis of `A`, if `{ b_j }` is an `F`-basis of `L`, + such that the family `{ a_i * b_j }` in `E` is `F`-linearly independent, + then `A` and `L` are linearly disjoint. + +Other main results: + +- `IntermediateField.LinearDisjoint.symm`, `IntermediateField.linearDisjoint_comm`: + linear disjointness is symmetric. + +- `IntermediateField.LinearDisjoint.rank_sup_of_isAlgebraic`, + `IntermediateField.LinearDisjoint.finrank_sup`: + if `A` and `B` are linearly disjoint, + then the rank of `A ⊔ B` is equal to the product of the rank of `A` and `B`. + + **TODO:** remove the algebraic assumptions (the proof becomes complicated). + +- `IntermediateField.LinearDisjoint.of_finrank_sup`: + conversely, if `A` and `B` are finite extensions, + such that rank of `A ⊔ B` is equal to the product of the rank of `A` and `B`, + then `A` and `B` are linearly disjoint. + +- `IntermediateField.LinearDisjoint.of_finrank_coprime`: + if the rank of `A` and `B` are coprime, + then `A` and `B` are linearly disjoint. + +- `IntermediateField.LinearDisjoint.inf_eq_bot`: + if `A` and `B` are linearly disjoint, then they are disjoint. + +## Tags + +linearly disjoint, linearly independent, tensor product + +-/ + +open scoped TensorProduct + +open Module IntermediateField + +noncomputable section + +universe u v w + +namespace IntermediateField + +variable {F : Type u} {E : Type v} [Field F] [Field E] [Algebra F E] + +variable (A B : IntermediateField F E) + +variable (L : Type w) [Field L] [Algebra F L] [Algebra L E] [IsScalarTower F L E] + +/-- If `A` is an intermediate field of `E / F`, and `E / L / F` is a field extension tower, +then `A` and `L` are linearly disjoint, if they are linearly disjoint as subalgebras of `E` +(`Subalgebra.LinearDisjoint`). -/ +protected abbrev LinearDisjoint : Prop := + A.toSubalgebra.LinearDisjoint (IsScalarTower.toAlgHom F L E).range + +theorem linearDisjoint_iff : + A.LinearDisjoint L ↔ A.toSubalgebra.LinearDisjoint (IsScalarTower.toAlgHom F L E).range := + Iff.rfl + +variable {A B L} + +/-- Two intermediate fields are linearly disjoint if and only if +they are linearly disjoint as subalgebras. -/ +theorem linearDisjoint_iff' : + A.LinearDisjoint B ↔ A.toSubalgebra.LinearDisjoint B.toSubalgebra := by + rw [linearDisjoint_iff] + congr! + ext; simp + +/-- Linear disjointness is symmetric. -/ +theorem LinearDisjoint.symm (H : A.LinearDisjoint B) : B.LinearDisjoint A := + linearDisjoint_iff'.2 (linearDisjoint_iff'.1 H).symm + +/-- Linear disjointness is symmetric. -/ +theorem linearDisjoint_comm : A.LinearDisjoint B ↔ B.LinearDisjoint A := + ⟨LinearDisjoint.symm, LinearDisjoint.symm⟩ + +namespace LinearDisjoint + +variable (A) in +theorem self_right : A.LinearDisjoint F := Subalgebra.LinearDisjoint.bot_right _ + +variable (A) in +theorem bot_right : A.LinearDisjoint (⊥ : IntermediateField F E) := + linearDisjoint_iff'.2 (Subalgebra.LinearDisjoint.bot_right _) + +variable (F E L) in +theorem bot_left : (⊥ : IntermediateField F E).LinearDisjoint L := + Subalgebra.LinearDisjoint.bot_left _ + +/-- If `A` and `L` are linearly disjoint, then any `F`-linearly independent family on `A` remains +linearly independent over `L`. -/ +theorem linearIndependent_left (H : A.LinearDisjoint L) + {ι : Type*} {a : ι → A} (ha : LinearIndependent F a) : LinearIndependent L (A.val ∘ a) := + (Subalgebra.LinearDisjoint.linearIndependent_left_of_flat H ha).map_of_injective_injective + (AlgEquiv.ofInjectiveField (IsScalarTower.toAlgHom F L E)) (AddMonoidHom.id E) + (by simp) (by simp) (fun _ _ ↦ by simp_rw [Algebra.smul_def]; rfl) + +/-- If there exists an `F`-basis of `A` which remains linearly independent over `L`, then +`A` and `L` are linearly disjoint. -/ +theorem of_basis_left {ι : Type*} (a : Basis ι F A) + (H : LinearIndependent L (A.val ∘ a)) : A.LinearDisjoint L := + Subalgebra.LinearDisjoint.of_basis_left _ _ a <| H.map_of_surjective_injective + (AlgEquiv.ofInjectiveField (IsScalarTower.toAlgHom F L E)) (AddMonoidHom.id E) + (AlgEquiv.surjective _) (by simp) (fun _ _ ↦ by simp_rw [Algebra.smul_def]; rfl) + +/-- If `A` and `B` are linearly disjoint, then any `F`-linearly independent family on `B` remains +linearly independent over `A`. -/ +theorem linearIndependent_right (H : A.LinearDisjoint B) + {ι : Type*} {b : ι → B} (hb : LinearIndependent F b) : LinearIndependent A (B.val ∘ b) := + (linearDisjoint_iff'.1 H).linearIndependent_right_of_flat hb + +/-- If there exists an `F`-basis of `B` which remains linearly independent over `A`, then +`A` and `B` are linearly disjoint. -/ +theorem of_basis_right {ι : Type*} (b : Basis ι F B) + (H : LinearIndependent A (B.val ∘ b)) : A.LinearDisjoint B := + linearDisjoint_iff'.2 (.of_basis_right _ _ b H) + +/-- If `A` and `L` are linearly disjoint, then any `F`-linearly independent family on `L` remains +linearly independent over `A`. -/ +theorem linearIndependent_right' (H : A.LinearDisjoint L) {ι : Type*} {b : ι → L} + (hb : LinearIndependent F b) : LinearIndependent A (algebraMap L E ∘ b) := by + apply Subalgebra.LinearDisjoint.linearIndependent_right_of_flat H <| hb.map' _ + (AlgEquiv.ofInjectiveField (IsScalarTower.toAlgHom F L E)).toLinearEquiv.ker + +/-- If there exists an `F`-basis of `L` which remains linearly independent over `A`, then +`A` and `L` are linearly disjoint. -/ +theorem of_basis_right' {ι : Type*} (b : Basis ι F L) + (H : LinearIndependent A (algebraMap L E ∘ b)) : A.LinearDisjoint L := + Subalgebra.LinearDisjoint.of_basis_right _ _ + (b.map (AlgEquiv.ofInjectiveField (IsScalarTower.toAlgHom F L E)).toLinearEquiv) H + +/-- If `A` and `B` are linearly disjoint, then for any `F`-linearly independent families +`{ u_i }`, `{ v_j }` of `A`, `B`, the products `{ u_i * v_j }` +are linearly independent over `F`. -/ +theorem linearIndependent_mul (H : A.LinearDisjoint B) {κ ι : Type*} {a : κ → A} {b : ι → B} + (ha : LinearIndependent F a) (hb : LinearIndependent F b) : + LinearIndependent F fun (i : κ × ι) ↦ (a i.1).1 * (b i.2).1 := + (linearDisjoint_iff'.1 H).linearIndependent_mul_of_flat_left ha hb + +/-- If `A` and `L` are linearly disjoint, then for any `F`-linearly independent families +`{ u_i }`, `{ v_j }` of `A`, `L`, the products `{ u_i * v_j }` +are linearly independent over `F`. -/ +theorem linearIndependent_mul' (H : A.LinearDisjoint L) {κ ι : Type*} {a : κ → A} {b : ι → L} + (ha : LinearIndependent F a) (hb : LinearIndependent F b) : + LinearIndependent F fun (i : κ × ι) ↦ (a i.1).1 * algebraMap L E (b i.2) := by + apply Subalgebra.LinearDisjoint.linearIndependent_mul_of_flat_left H ha <| hb.map' _ + (AlgEquiv.ofInjectiveField (IsScalarTower.toAlgHom F L E)).toLinearEquiv.ker + +/-- If there are `F`-bases `{ u_i }`, `{ v_j }` of `A`, `B`, such that the products +`{ u_i * v_j }` are linearly independent over `F`, then `A` and `B` are linearly disjoint. -/ +theorem of_basis_mul {κ ι : Type*} (a : Basis κ F A) (b : Basis ι F B) + (H : LinearIndependent F fun (i : κ × ι) ↦ (a i.1).1 * (b i.2).1) : A.LinearDisjoint B := + linearDisjoint_iff'.2 (.of_basis_mul _ _ a b H) + +/-- If there are `F`-bases `{ u_i }`, `{ v_j }` of `A`, `L`, such that the products +`{ u_i * v_j }` are linearly independent over `F`, then `A` and `L` are linearly disjoint. -/ +theorem of_basis_mul' {κ ι : Type*} (a : Basis κ F A) (b : Basis ι F L) + (H : LinearIndependent F fun (i : κ × ι) ↦ (a i.1).1 * algebraMap L E (b i.2)) : + A.LinearDisjoint L := + Subalgebra.LinearDisjoint.of_basis_mul _ _ a + (b.map (AlgEquiv.ofInjectiveField (IsScalarTower.toAlgHom F L E)).toLinearEquiv) H + +theorem of_le_left {A' : IntermediateField F E} (H : A.LinearDisjoint L) + (h : A' ≤ A) : A'.LinearDisjoint L := + Subalgebra.LinearDisjoint.of_le_left_of_flat H h + +theorem of_le_right {B' : IntermediateField F E} (H : A.LinearDisjoint B) + (h : B' ≤ B) : A.LinearDisjoint B' := + linearDisjoint_iff'.2 ((linearDisjoint_iff'.1 H).of_le_right_of_flat h) + +/-- Similar to `IntermediateField.LinearDisjoint.of_le_right` but this is for abstract fields. -/ +theorem of_le_right' (H : A.LinearDisjoint L) (L' : Type*) [Field L'] + [Algebra F L'] [Algebra L' L] [IsScalarTower F L' L] + [Algebra L' E] [IsScalarTower F L' E] [IsScalarTower L' L E] : A.LinearDisjoint L' := by + refine Subalgebra.LinearDisjoint.of_le_right_of_flat H ?_ + convert AlgHom.range_comp_le_range (IsScalarTower.toAlgHom F L' L) (IsScalarTower.toAlgHom F L E) + ext; exact IsScalarTower.algebraMap_apply L' L E _ + +/-- If `A` and `B` are linearly disjoint, `A'` and `B'` are contained in `A` and `B`, +respectively, then `A'` and `B'` are also linearly disjoint. -/ +theorem of_le {A' B' : IntermediateField F E} (H : A.LinearDisjoint B) + (hA : A' ≤ A) (hB : B' ≤ B) : A'.LinearDisjoint B' := + H.of_le_left hA |>.of_le_right hB + +/-- Similar to `IntermediateField.LinearDisjoint.of_le` but this is for abstract fields. -/ +theorem of_le' {A' : IntermediateField F E} (H : A.LinearDisjoint L) + (hA : A' ≤ A) (L' : Type*) [Field L'] + [Algebra F L'] [Algebra L' L] [IsScalarTower F L' L] + [Algebra L' E] [IsScalarTower F L' E] [IsScalarTower L' L E] : A'.LinearDisjoint L' := + H.of_le_left hA |>.of_le_right' L' + +/-- If `A` and `B` are linearly disjoint over `F`, then their intersection is equal to `F`. -/ +theorem inf_eq_bot (H : A.LinearDisjoint B) : + A ⊓ B = ⊥ := toSubalgebra_injective (linearDisjoint_iff'.1 H).inf_eq_bot + +/-- If `A` and `A` itself are linearly disjoint over `F`, then it is equal to `F`. -/ +theorem eq_bot_of_self (H : A.LinearDisjoint A) : A = ⊥ := + inf_idem A ▸ H.inf_eq_bot + +/-- If `A` and `B` are linearly disjoint over `F`, and at least one them are algebraic, then the +rank of `A ⊔ B` is equal to the product of that of `A` and `B`. Note that this result is +also true without algebraic assumption, but the proof becomes very complicated. -/ +theorem rank_sup_of_isAlgebraic (H : A.LinearDisjoint B) + (halg : Algebra.IsAlgebraic F A ∨ Algebra.IsAlgebraic F B) : + Module.rank F ↥(A ⊔ B) = Module.rank F A * Module.rank F B := + have h := le_sup_toSubalgebra A B + (rank_sup_le_of_isAlgebraic A B halg).antisymm <| + (linearDisjoint_iff'.1 H).rank_sup_of_free.ge.trans <| + (Subalgebra.inclusion h).toLinearMap.rank_le_of_injective (Subalgebra.inclusion_injective h) + +/-- If `A` and `B` are linearly disjoint over `F`, then the `Module.finrank` of +`A ⊔ B` is equal to the product of that of `A` and `B`. -/ +theorem finrank_sup (H : A.LinearDisjoint B) : finrank F ↥(A ⊔ B) = finrank F A * finrank F B := by + by_cases h : FiniteDimensional F A + · simpa only [map_mul] using + congr(Cardinal.toNat $(H.rank_sup_of_isAlgebraic (.inl inferInstance))) + rw [FiniteDimensional, ← rank_lt_aleph0_iff, not_lt] at h + have := LinearMap.rank_le_of_injective _ <| Submodule.inclusion_injective <| + show Subalgebra.toSubmodule A.toSubalgebra ≤ Subalgebra.toSubmodule (A ⊔ B).toSubalgebra by simp + rw [show finrank F A = 0 from Cardinal.toNat_apply_of_aleph0_le h, + show finrank F ↥(A ⊔ B) = 0 from Cardinal.toNat_apply_of_aleph0_le (h.trans this), zero_mul] + +/-- If `A` and `B` are finite extensions of `F`, +such that rank of `A ⊔ B` is equal to the product of the rank of `A` and `B`, +then `A` and `B` are linearly disjoint. -/ +theorem of_finrank_sup [FiniteDimensional F A] [FiniteDimensional F B] + (H : finrank F ↥(A ⊔ B) = finrank F A * finrank F B) : A.LinearDisjoint B := + linearDisjoint_iff'.2 <| .of_finrank_sup_of_free (by rwa [← sup_toSubalgebra_of_left]) + +/-- If `A` and `L` have coprime degree over `F`, then they are linearly disjoint. -/ +theorem of_finrank_coprime (H : (finrank F A).Coprime (finrank F L)) : A.LinearDisjoint L := + letI : Field (AlgHom.range (IsScalarTower.toAlgHom F L E)) := + inferInstanceAs <| Field (AlgHom.fieldRange (IsScalarTower.toAlgHom F L E)) + letI : Field A.toSubalgebra := inferInstanceAs <| Field A + Subalgebra.LinearDisjoint.of_finrank_coprime_of_free <| by + rwa [(AlgEquiv.ofInjectiveField (IsScalarTower.toAlgHom F L E)).toLinearEquiv.finrank_eq] at H + +end LinearDisjoint + +end IntermediateField diff --git a/Mathlib/LinearAlgebra/LinearDisjoint.lean b/Mathlib/LinearAlgebra/LinearDisjoint.lean index 18ed93e9d07ed..260ea22120c2e 100644 --- a/Mathlib/LinearAlgebra/LinearDisjoint.lean +++ b/Mathlib/LinearAlgebra/LinearDisjoint.lean @@ -25,7 +25,7 @@ Let `M` and `N` be `R`-submodules in `S` (`Submodule R S`). `M.LinearDisjoint N`), if the natural `R`-linear map `M ⊗[R] N →ₗ[R] S` (`Submodule.mulMap M N`) induced by the multiplication in `S` is injective. -The following is the first equivalent characterization of linearly disjointness: +The following is the first equivalent characterization of linear disjointness: - `Submodule.LinearDisjoint.linearIndependent_left_of_flat`: if `M` and `N` are linearly disjoint, if `N` is a flat `R`-module, then for any family of @@ -49,7 +49,7 @@ Dually, we have: conversely, if `{ n_i }` is an `R`-basis of `N`, which is also `M`-linearly independent, then `M` and `N` are linearly disjoint. -The following is the second equivalent characterization of linearly disjointness: +The following is the second equivalent characterization of linear disjointness: - `Submodule.LinearDisjoint.linearIndependent_mul_of_flat`: if `M` and `N` are linearly disjoint, if one of `M` and `N` is flat, then for any family of @@ -64,15 +64,15 @@ The following is the second equivalent characterization of linearly disjointness ## Other main results -- `Submodule.LinearDisjoint.symm_of_commute`, `Submodule.linearDisjoint_symm_of_commute`: - linearly disjoint is symmetric under some commutative conditions. +- `Submodule.LinearDisjoint.symm_of_commute`, `Submodule.linearDisjoint_comm_of_commute`: + linear disjointness is symmetric under some commutative conditions. - `Submodule.linearDisjoint_op`: - linearly disjoint is preserved by taking multiplicative opposite. + linear disjointness is preserved by taking multiplicative opposite. - `Submodule.LinearDisjoint.of_le_left_of_flat`, `Submodule.LinearDisjoint.of_le_right_of_flat`, `Submodule.LinearDisjoint.of_le_of_flat_left`, `Submodule.LinearDisjoint.of_le_of_flat_right`: - linearly disjoint is preserved by taking submodules under some flatness conditions. + linear disjointness is preserved by taking submodules under some flatness conditions. - `Submodule.LinearDisjoint.of_linearDisjoint_fg_left`, `Submodule.LinearDisjoint.of_linearDisjoint_fg_right`, @@ -162,7 +162,7 @@ theorem LinearDisjoint.of_subsingleton [Subsingleton R] : M.LinearDisjoint N := haveI : Subsingleton S := Module.subsingleton R S exact ⟨Function.injective_of_subsingleton _⟩ -/-- Linearly disjoint is preserved by taking multiplicative opposite. -/ +/-- Linear disjointness is preserved by taking multiplicative opposite. -/ theorem linearDisjoint_op : M.LinearDisjoint N ↔ (equivOpposite.symm (MulOpposite.op N)).LinearDisjoint (equivOpposite.symm (MulOpposite.op M)) := by @@ -171,14 +171,14 @@ theorem linearDisjoint_op : alias ⟨LinearDisjoint.op, LinearDisjoint.of_op⟩ := linearDisjoint_op -/-- Linearly disjoint is symmetric if elements in the module commute. -/ +/-- Linear disjointness is symmetric if elements in the module commute. -/ theorem LinearDisjoint.symm_of_commute (H : M.LinearDisjoint N) (hc : ∀ (m : M) (n : N), Commute m.1 n.1) : N.LinearDisjoint M := by rw [linearDisjoint_iff, mulMap_comm_of_commute M N hc] exact ((TensorProduct.comm R N M).toEquiv.injective_comp _).2 H.injective -/-- Linearly disjoint is symmetric if elements in the module commute. -/ -theorem linearDisjoint_symm_of_commute +/-- Linear disjointness is symmetric if elements in the module commute. -/ +theorem linearDisjoint_comm_of_commute (hc : ∀ (m : M) (n : N), Commute m.1 n.1) : M.LinearDisjoint N ↔ N.LinearDisjoint M := ⟨fun H ↦ H.symm_of_commute hc, fun H ↦ H.symm_of_commute fun _ _ ↦ (hc _ _).symm⟩ @@ -284,12 +284,12 @@ variable [CommSemiring R] [CommSemiring S] [Algebra R S] variable {M N : Submodule R S} -/-- Linearly disjoint is symmetric in a commutative ring. -/ +/-- Linear disjointness is symmetric in a commutative ring. -/ theorem LinearDisjoint.symm (H : M.LinearDisjoint N) : N.LinearDisjoint M := H.symm_of_commute fun _ _ ↦ mul_comm _ _ -/-- Linearly disjoint is symmetric in a commutative ring. -/ -theorem linearDisjoint_symm : M.LinearDisjoint N ↔ N.LinearDisjoint M := +/-- Linear disjointness is symmetric in a commutative ring. -/ +theorem linearDisjoint_comm : M.LinearDisjoint N ↔ N.LinearDisjoint M := ⟨LinearDisjoint.symm, LinearDisjoint.symm⟩ end CommSemiring diff --git a/Mathlib/RingTheory/LinearDisjoint.lean b/Mathlib/RingTheory/LinearDisjoint.lean index 11a726e495c0b..b03dac78930c7 100644 --- a/Mathlib/RingTheory/LinearDisjoint.lean +++ b/Mathlib/RingTheory/LinearDisjoint.lean @@ -32,7 +32,7 @@ See the file `Mathlib/LinearAlgebra/LinearDisjoint.lean` for details. ## Main results -### Equivalent characterization of linearly disjointness +### Equivalent characterization of linear disjointness - `Subalgebra.LinearDisjoint.linearIndependent_left_of_flat`: if `A` and `B` are linearly disjoint, and if `B` is a flat `R`-module, then for any family of @@ -63,8 +63,8 @@ See the file `Mathlib/LinearAlgebra/LinearDisjoint.lean` for details. ### Other main results -- `Subalgebra.LinearDisjoint.symm_of_commute`, `Subalgebra.linearDisjoint_symm_of_commute`: - linearly disjoint is symmetric under some commutative conditions. +- `Subalgebra.LinearDisjoint.symm_of_commute`, `Subalgebra.linearDisjoint_comm_of_commute`: + linear disjointness is symmetric under some commutative conditions. - `Subalgebra.LinearDisjoint.bot_left`, `Subalgebra.LinearDisjoint.bot_right`: the image of `R` in `S` is linearly disjoint with any other subalgebras. @@ -104,7 +104,7 @@ linearly disjoint, linearly independent, tensor product -/ -open scoped Classical TensorProduct +open scoped TensorProduct noncomputable section @@ -133,13 +133,13 @@ variable {A B} theorem LinearDisjoint.of_subsingleton [Subsingleton R] : A.LinearDisjoint B := Submodule.LinearDisjoint.of_subsingleton -/-- Linearly disjoint is symmetric if elements in the module commute. -/ +/-- Linear disjointness is symmetric if elements in the module commute. -/ theorem LinearDisjoint.symm_of_commute (H : A.LinearDisjoint B) (hc : ∀ (a : A) (b : B), Commute a.1 b.1) : B.LinearDisjoint A := Submodule.LinearDisjoint.symm_of_commute H hc -/-- Linearly disjoint is symmetric if elements in the module commute. -/ -theorem linearDisjoint_symm_of_commute +/-- Linear disjointness is symmetric if elements in the module commute. -/ +theorem linearDisjoint_comm_of_commute (hc : ∀ (a : A) (b : B), Commute a.1 b.1) : A.LinearDisjoint B ↔ B.LinearDisjoint A := ⟨fun H ↦ H.symm_of_commute hc, fun H ↦ H.symm_of_commute fun _ _ ↦ (hc _ _).symm⟩ @@ -167,12 +167,12 @@ variable [CommSemiring R] [CommSemiring S] [Algebra R S] variable {A B : Subalgebra R S} -/-- Linearly disjoint is symmetric in a commutative ring. -/ +/-- Linear disjointness is symmetric in a commutative ring. -/ theorem LinearDisjoint.symm (H : A.LinearDisjoint B) : B.LinearDisjoint A := H.symm_of_commute fun _ _ ↦ mul_comm _ _ -/-- Linearly disjoint is symmetric in a commutative ring. -/ -theorem linearDisjoint_symm : A.LinearDisjoint B ↔ B.LinearDisjoint A := +/-- Linear disjointness is symmetric in a commutative ring. -/ +theorem linearDisjoint_comm : A.LinearDisjoint B ↔ B.LinearDisjoint A := ⟨LinearDisjoint.symm, LinearDisjoint.symm⟩ namespace LinearDisjoint From f899f09c0d59f45ba460b323c102500e38b34f26 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Wed, 27 Nov 2024 09:04:45 +0000 Subject: [PATCH 339/829] chore: avoid importing `ContDiff.Defs` in `FDeriv.Analytic` (#19374) For this, move the results needing `ContDiff` to two new files. The reason of this change is that I will import `FDeriv.Analytic` in `ContDiff.Defs` in #17152 --- Mathlib.lean | 2 + Mathlib/Analysis/Analytic/Within.lean | 29 +++++-- .../Analysis/Calculus/ContDiff/Analytic.lean | 55 +++++++++++++ .../Calculus/ContDiff/CPolynomial.lean | 63 +++++++++++++++ .../Analysis/Calculus/FDeriv/Analytic.lean | 79 +++---------------- .../Fourier/FourierTransformDeriv.lean | 3 +- .../Analysis/SpecialFunctions/ExpDeriv.lean | 2 + .../Geometry/Manifold/AnalyticManifold.lean | 2 +- 8 files changed, 155 insertions(+), 80 deletions(-) create mode 100644 Mathlib/Analysis/Calculus/ContDiff/Analytic.lean create mode 100644 Mathlib/Analysis/Calculus/ContDiff/CPolynomial.lean diff --git a/Mathlib.lean b/Mathlib.lean index 26f7426ffef46..21358bdd55951 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1120,8 +1120,10 @@ import Mathlib.Analysis.Calculus.BumpFunction.InnerProduct import Mathlib.Analysis.Calculus.BumpFunction.Normed import Mathlib.Analysis.Calculus.Conformal.InnerProduct import Mathlib.Analysis.Calculus.Conformal.NormedSpace +import Mathlib.Analysis.Calculus.ContDiff.Analytic import Mathlib.Analysis.Calculus.ContDiff.Basic import Mathlib.Analysis.Calculus.ContDiff.Bounds +import Mathlib.Analysis.Calculus.ContDiff.CPolynomial import Mathlib.Analysis.Calculus.ContDiff.Defs import Mathlib.Analysis.Calculus.ContDiff.FTaylorSeries import Mathlib.Analysis.Calculus.ContDiff.FaaDiBruno diff --git a/Mathlib/Analysis/Analytic/Within.lean b/Mathlib/Analysis/Analytic/Within.lean index 67e0db70351e1..fd9bd854f032d 100644 --- a/Mathlib/Analysis/Analytic/Within.lean +++ b/Mathlib/Analysis/Analytic/Within.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Geoffrey Irving -/ import Mathlib.Analysis.Analytic.Constructions +import Mathlib.Analysis.Analytic.ChangeOrigin /-! # Properties of analyticity restricted to a set @@ -40,8 +41,8 @@ lemma analyticWithinAt_of_singleton_mem {f : E → F} {s : Set E} {x : E} (h : { AnalyticWithinAt 𝕜 f s x := by rcases mem_nhdsWithin.mp h with ⟨t, ot, xt, st⟩ rcases Metric.mem_nhds_iff.mp (ot.mem_nhds xt) with ⟨r, r0, rt⟩ - exact ⟨constFormalMultilinearSeries 𝕜 E (f x), .ofReal r, { - r_le := by simp only [FormalMultilinearSeries.constFormalMultilinearSeries_radius, le_top] + exact ⟨constFormalMultilinearSeries 𝕜 E (f x), .ofReal r, + { r_le := by simp only [FormalMultilinearSeries.constFormalMultilinearSeries_radius, le_top] r_pos := by positivity hasSum := by intro y ys yr @@ -63,10 +64,10 @@ lemma analyticOn_of_locally_analyticOn {f : E → F} {s : Set E} rcases h x m with ⟨u, ou, xu, fu⟩ rcases Metric.mem_nhds_iff.mp (ou.mem_nhds xu) with ⟨r, r0, ru⟩ rcases fu x ⟨m, xu⟩ with ⟨p, t, fp⟩ - exact ⟨p, min (.ofReal r) t, { - r_pos := lt_min (by positivity) fp.r_pos - r_le := min_le_of_right_le fp.r_le - hasSum := by + exact ⟨p, min (.ofReal r) t, + { r_pos := lt_min (by positivity) fp.r_pos + r_le := min_le_of_right_le fp.r_le + hasSum := by intro y ys yr simp only [EMetric.mem_ball, lt_min_iff, edist_lt_ofReal, dist_zero_right] at yr apply fp.hasSum @@ -88,8 +89,8 @@ lemma IsOpen.analyticOn_iff_analyticOnNhd {f : E → F} {s : Set E} (hs : IsOpen intro hf x m rcases Metric.mem_nhds_iff.mp (hs.mem_nhds m) with ⟨r, r0, rs⟩ rcases hf x m with ⟨p, t, fp⟩ - exact ⟨p, min (.ofReal r) t, { - r_pos := lt_min (by positivity) fp.r_pos + exact ⟨p, min (.ofReal r) t, + { r_pos := lt_min (by positivity) fp.r_pos r_le := min_le_of_right_le fp.r_le hasSum := by intro y ym @@ -200,3 +201,15 @@ lemma analyticWithinAt_iff_exists_analyticAt' [CompleteSpace F] {f : E → F} {s exact ⟨g, by filter_upwards [self_mem_nhdsWithin] using hf, hg⟩ alias ⟨AnalyticWithinAt.exists_analyticAt, _⟩ := analyticWithinAt_iff_exists_analyticAt' + +lemma AnalyticWithinAt.exists_mem_nhdsWithin_analyticOn + [CompleteSpace F] {f : E → F} {s : Set E} {x : E} (h : AnalyticWithinAt 𝕜 f s x) : + ∃ u ∈ 𝓝[insert x s] x, AnalyticOn 𝕜 f u := by + obtain ⟨g, -, h'g, hg⟩ : ∃ g, f x = g x ∧ EqOn f g (insert x s) ∧ AnalyticAt 𝕜 g x := + h.exists_analyticAt + let u := insert x s ∩ {y | AnalyticAt 𝕜 g y} + refine ⟨u, ?_, ?_⟩ + · exact inter_mem_nhdsWithin _ ((isOpen_analyticAt 𝕜 g).mem_nhds hg) + · intro y hy + have : AnalyticWithinAt 𝕜 g u y := hy.2.analyticWithinAt + exact this.congr (h'g.mono (inter_subset_left)) (h'g (inter_subset_left hy)) diff --git a/Mathlib/Analysis/Calculus/ContDiff/Analytic.lean b/Mathlib/Analysis/Calculus/ContDiff/Analytic.lean new file mode 100644 index 0000000000000..8ba3cda66b843 --- /dev/null +++ b/Mathlib/Analysis/Calculus/ContDiff/Analytic.lean @@ -0,0 +1,55 @@ +/- +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.Analysis.Calculus.ContDiff.Defs +import Mathlib.Analysis.Calculus.FDeriv.Analytic + +/-! +# Analytic functions are `C^∞`. +-/ + +open Set ContDiff + +variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] + {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] + {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] + {f : E → F} {s : Set E} {x : E} {n : WithTop ℕ∞} + +/-- An analytic function is infinitely differentiable. -/ +protected theorem AnalyticOnNhd.contDiffOn [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f s) : + ContDiffOn 𝕜 n f s := by + let t := { x | AnalyticAt 𝕜 f x } + suffices ContDiffOn 𝕜 ω f t from (this.of_le le_top).mono h + rw [← contDiffOn_infty_iff_contDiffOn_omega] + have H : AnalyticOnNhd 𝕜 f t := fun _x hx ↦ hx + have t_open : IsOpen t := isOpen_analyticAt 𝕜 f + exact contDiffOn_of_continuousOn_differentiableOn + (fun m _ ↦ (H.iteratedFDeriv m).continuousOn.congr + fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) + (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) : + 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 + exact hf.contDiffOn.contDiffAt hs + +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 + +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 diff --git a/Mathlib/Analysis/Calculus/ContDiff/CPolynomial.lean b/Mathlib/Analysis/Calculus/ContDiff/CPolynomial.lean new file mode 100644 index 0000000000000..3a66c7c3d570a --- /dev/null +++ b/Mathlib/Analysis/Calculus/ContDiff/CPolynomial.lean @@ -0,0 +1,63 @@ +/- +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.Analysis.Calculus.FDeriv.Analytic +import Mathlib.Analysis.Calculus.ContDiff.Defs + +/-! +# Higher smoothness of continuously polynomial functions + +We prove that continuously polynomial functions are `C^∞`. In particular, this is the case +of continuous multilinear maps. +-/ + +open Filter Asymptotics + +open scoped ENNReal ContDiff + +universe u v + +variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] +variable {E : Type u} [NormedAddCommGroup E] [NormedSpace 𝕜 E] +variable {F : Type v} [NormedAddCommGroup F] [NormedSpace 𝕜 F] + +section fderiv + +variable {p : FormalMultilinearSeries 𝕜 E F} {r : ℝ≥0∞} {n : ℕ} +variable {f : E → F} {x : E} {s : Set E} + +/-- A polynomial function is infinitely differentiable. -/ +theorem CPolynomialOn.contDiffOn (h : CPolynomialOn 𝕜 f s) {n : WithTop ℕ∞} : + ContDiffOn 𝕜 n f s := by + let t := { x | CPolynomialAt 𝕜 f x } + suffices ContDiffOn 𝕜 ω f t from (this.of_le le_top).mono h + rw [← contDiffOn_infty_iff_contDiffOn_omega] + have H : CPolynomialOn 𝕜 f t := fun _x hx ↦ hx + have t_open : IsOpen t := isOpen_cPolynomialAt 𝕜 f + exact contDiffOn_of_continuousOn_differentiableOn + (fun m _ ↦ (H.iteratedFDeriv m).continuousOn.congr + fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) + (fun m _ ↦ (H.iteratedFDeriv m).analyticOnNhd.differentiableOn.congr + fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) + +theorem CPolynomialAt.contDiffAt (h : CPolynomialAt 𝕜 f x) {n : WithTop ℕ∞} : + ContDiffAt 𝕜 n f x := + let ⟨_, hs, hf⟩ := h.exists_mem_nhds_cPolynomialOn + hf.contDiffOn.contDiffAt hs + +end fderiv + +namespace ContinuousMultilinearMap + +variable {ι : Type*} {E : ι → Type*} [∀ i, NormedAddCommGroup (E i)] [∀ i, NormedSpace 𝕜 (E i)] + [Fintype ι] (f : ContinuousMultilinearMap 𝕜 E F) {n : WithTop ℕ∞} {x : Π i, E i} + +open FormalMultilinearSeries + +lemma contDiffAt : ContDiffAt 𝕜 n f x := f.cpolynomialAt.contDiffAt + +lemma contDiff : ContDiff 𝕜 n f := contDiff_iff_contDiffAt.mpr (fun _ ↦ f.contDiffAt) + +end ContinuousMultilinearMap diff --git a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean index 9141ffab19f41..48cb70ce773fd 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean @@ -7,7 +7,7 @@ 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.ContDiff.FTaylorSeries import Mathlib.Analysis.Calculus.FDeriv.Add import Mathlib.Analysis.Calculus.FDeriv.Prod import Mathlib.Analysis.Normed.Module.Completion @@ -63,7 +63,7 @@ differentiability at points in a neighborhood of `s`. Therefore, the theorem tha open Filter Asymptotics Set -open scoped ENNReal Topology ContDiff +open scoped ENNReal Topology universe u v @@ -192,6 +192,7 @@ theorem HasFPowerSeriesOnBall.fderiv_eq [CompleteSpace F] (h : HasFPowerSeriesOn 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. -/ protected theorem HasFPowerSeriesOnBall.fderiv [CompleteSpace F] (h : HasFPowerSeriesOnBall f p x r) : HasFPowerSeriesOnBall (fderiv 𝕜 f) p.derivSeries x r := by @@ -238,7 +239,7 @@ protected theorem AnalyticOnNhd.fderiv [CompleteSpace F] (h : AnalyticOnNhd 𝕜 alias AnalyticOn.fderiv := AnalyticOnNhd.fderiv /-- 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 +`AnalyticOnNhd.iteratedFDeriv_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 @@ -269,44 +270,6 @@ lemma AnalyticOnNhd.hasFTaylorSeriesUpToOn [CompleteSpace F] · apply (DifferentiableAt.continuousAt (𝕜 := 𝕜) ?_).continuousWithinAt exact (h.iteratedFDeriv m x hx).differentiableAt -/-- An analytic function is infinitely differentiable. -/ -protected theorem AnalyticOnNhd.contDiffOn [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f s) - {n : WithTop ℕ∞} : ContDiffOn 𝕜 n f s := by - suffices ContDiffOn 𝕜 ω f s from this.of_le le_top - rw [← contDiffOn_infty_iff_contDiffOn_omega] - let t := { x | AnalyticAt 𝕜 f x } - suffices ContDiffOn 𝕜 ∞ f t from this.mono h - have H : AnalyticOnNhd 𝕜 f t := fun _x hx ↦ hx - have t_open : IsOpen t := isOpen_analyticAt 𝕜 f - exact contDiffOn_of_continuousOn_differentiableOn - (fun m _ ↦ (H.iteratedFDeriv m).continuousOn.congr - fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) - (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 : WithTop ℕ∞} : - ContDiff 𝕜 n f := by - rw [← contDiffOn_univ] - exact h.contDiffOn - -theorem AnalyticAt.contDiffAt [CompleteSpace F] (h : AnalyticAt 𝕜 f x) {n : WithTop ℕ∞} : - ContDiffAt 𝕜 n f x := by - obtain ⟨s, hs, hf⟩ := h.exists_mem_nhds_analyticOnNhd - exact hf.contDiffOn.contDiffAt hs - -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 - -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 : WithTop ℕ∞) (h : AnalyticWithinAt 𝕜 f s x) : ∃ u ∈ 𝓝[insert x s] x, ∃ (p : E → FormalMultilinearSeries 𝕜 E F), @@ -384,7 +347,7 @@ protected theorem AnalyticOn.iteratedFDerivWithin (h : AnalyticOn 𝕜 f s) apply AnalyticOnNhd.comp_analyticOn _ (IH.fderivWithin hu) (mapsTo_univ _ _) apply LinearIsometryEquiv.analyticOnNhd -lemma AnalyticOn.hasFTaylorSeriesUpToOn {n : WithTop ℕ∞} +protected lemma AnalyticOn.hasFTaylorSeriesUpToOn {n : WithTop ℕ∞} (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 ↦ ?_⟩ @@ -546,26 +509,6 @@ theorem CPolynomialOn.iteratedFDeriv (h : CPolynomialOn 𝕜 f s) (n : ℕ) : case g => exact ↑(continuousMultilinearCurryLeftEquiv 𝕜 (fun _ : Fin (n + 1) ↦ E) F).symm simp -/-- A polynomial function is infinitely differentiable. -/ -theorem CPolynomialOn.contDiffOn (h : CPolynomialOn 𝕜 f s) {n : WithTop ℕ∞} : - ContDiffOn 𝕜 n f s := by - suffices ContDiffOn 𝕜 ω f s from this.of_le le_top - let t := { x | CPolynomialAt 𝕜 f x } - suffices ContDiffOn 𝕜 ω f t from this.mono h - rw [← contDiffOn_infty_iff_contDiffOn_omega] - have H : CPolynomialOn 𝕜 f t := fun _x hx ↦ hx - have t_open : IsOpen t := isOpen_cPolynomialAt 𝕜 f - exact contDiffOn_of_continuousOn_differentiableOn - (fun m _ ↦ (H.iteratedFDeriv m).continuousOn.congr - fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) - (fun m _ ↦ (H.iteratedFDeriv m).analyticOnNhd.differentiableOn.congr - fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) - -theorem CPolynomialAt.contDiffAt (h : CPolynomialAt 𝕜 f x) {n : WithTop ℕ∞} : - ContDiffAt 𝕜 n f x := - let ⟨_, hs, hf⟩ := h.exists_mem_nhds_cPolynomialOn - hf.contDiffOn.contDiffAt hs - end fderiv section deriv @@ -750,10 +693,6 @@ lemma cPolynomialAt : CPolynomialAt 𝕜 f x := lemma cPolyomialOn : CPolynomialOn 𝕜 f ⊤ := fun x _ ↦ f.cPolynomialAt x -lemma contDiffAt : ContDiffAt 𝕜 n f x := (f.cPolynomialAt x).contDiffAt - -lemma contDiff : ContDiff 𝕜 n f := contDiff_iff_contDiffAt.mpr f.contDiffAt - end ContinuousMultilinearMap namespace FormalMultilinearSeries @@ -797,8 +736,8 @@ private theorem factorial_smul' {n : ℕ} : ∀ {F : Type max u v} [NormedAddCom n ! • p n (fun _ ↦ y) = iteratedFDeriv 𝕜 n f x (fun _ ↦ y) := by induction n with | zero => _ | succ n ih => _ <;> intro F _ _ _ p f h · rw [factorial_zero, one_smul, h.iteratedFDeriv_zero_apply_diag] - · rw [factorial_succ, mul_comm, mul_smul, ← derivSeries_apply_diag, ← smul_apply, - ih h.fderiv, iteratedFDeriv_succ_apply_right] + · rw [factorial_succ, mul_comm, mul_smul, ← derivSeries_apply_diag, + ← ContinuousLinearMap.smul_apply, ih h.fderiv, iteratedFDeriv_succ_apply_right] rfl variable [CompleteSpace F] @@ -808,8 +747,8 @@ theorem factorial_smul (n : ℕ) : n ! • p n (fun _ ↦ y) = iteratedFDeriv 𝕜 n f x (fun _ ↦ y) := by cases n · rw [factorial_zero, one_smul, h.iteratedFDeriv_zero_apply_diag] - · rw [factorial_succ, mul_comm, mul_smul, ← derivSeries_apply_diag, ← smul_apply, - factorial_smul' _ h.fderiv, iteratedFDeriv_succ_apply_right] + · rw [factorial_succ, mul_comm, mul_smul, ← derivSeries_apply_diag, + ← ContinuousLinearMap.smul_apply, factorial_smul' _ h.fderiv, iteratedFDeriv_succ_apply_right] rfl theorem hasSum_iteratedFDeriv [CharZero 𝕜] {y : E} (hy : y ∈ EMetric.ball 0 r) : diff --git a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean index 0151b54688137..dcf554692ab17 100644 --- a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean +++ b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Alex Kontorovich, David Loeffler, Heather Macbeth, Sébastien Gouëzel -/ import Mathlib.Analysis.Calculus.ParametricIntegral +import Mathlib.Analysis.Calculus.ContDiff.CPolynomial import Mathlib.Analysis.Fourier.AddCircle import Mathlib.Analysis.Fourier.FourierTransform import Mathlib.Analysis.Calculus.FDeriv.Analytic @@ -369,7 +370,7 @@ lemma norm_iteratedFDeriv_fourierPowSMulRight have I₂ m : ‖iteratedFDeriv ℝ m (T ∘ (ContinuousLinearMap.pi (fun (_ : Fin n) ↦ L))) v‖ ≤ (n.descFactorial m * 1 * (‖L‖ * ‖v‖) ^ (n - m)) * ‖L‖ ^ m := by rw [ContinuousLinearMap.iteratedFDeriv_comp_right _ (ContinuousMultilinearMap.contDiff _) - _ le_top] + _ (mod_cast le_top)] apply (norm_compContinuousLinearMap_le _ _).trans simp only [Finset.prod_const, Finset.card_fin] gcongr diff --git a/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean b/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean index 846339ecb2f00..7a9ffe7ff6f72 100644 --- a/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean @@ -4,7 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne -/ import Mathlib.Analysis.Complex.RealDeriv +import Mathlib.Analysis.Calculus.ContDiff.Analytic import Mathlib.Analysis.Calculus.ContDiff.RCLike +import Mathlib.Analysis.Calculus.FDeriv.Analytic import Mathlib.Analysis.Calculus.IteratedDeriv.Lemmas import Mathlib.Analysis.SpecialFunctions.Exponential diff --git a/Mathlib/Geometry/Manifold/AnalyticManifold.lean b/Mathlib/Geometry/Manifold/AnalyticManifold.lean index 9a1ab30e1daa9..ea26292a58a0d 100644 --- a/Mathlib/Geometry/Manifold/AnalyticManifold.lean +++ b/Mathlib/Geometry/Manifold/AnalyticManifold.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Lee, Geoffrey Irving -/ import Mathlib.Analysis.Analytic.Constructions -import Mathlib.Analysis.Calculus.FDeriv.Analytic +import Mathlib.Analysis.Calculus.ContDiff.Analytic import Mathlib.Geometry.Manifold.SmoothManifoldWithCorners /-! From e10334bd8b2298ee04df07fbd1896ffb60e143b6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 27 Nov 2024 09:04:46 +0000 Subject: [PATCH 340/829] =?UTF-8?q?chore:=20rename=20`div=5Fdiv=5Fcancel'`?= =?UTF-8?q?=20to=20`div=5Fdiv=5Fcancel=E2=82=80`=20(#19459)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit and make it simp --- Mathlib/Algebra/GroupWithZero/Units/Basic.lean | 4 +++- Mathlib/Algebra/Order/Field/Basic.lean | 2 +- Mathlib/Analysis/Complex/Liouville.lean | 2 +- Mathlib/Analysis/Complex/PhragmenLindelof.lean | 2 +- Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean | 2 +- 5 files changed, 7 insertions(+), 5 deletions(-) diff --git a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean index be6aebca0a776..6dcca5a5b9844 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean @@ -451,7 +451,9 @@ lemma div_eq_div_iff_div_eq_div' (hb : b ≠ 0) (hc : c ≠ 0) : a / b = c / d conv_rhs => rw [← mul_left_inj' hc, div_mul_cancel₀ _ hc] rw [mul_comm _ c, div_mul_eq_mul_div, mul_div_assoc] -lemma div_div_cancel' (ha : a ≠ 0) : a / (a / b) = b := ha.isUnit.div_div_cancel +@[simp] lemma div_div_cancel₀ (ha : a ≠ 0) : a / (a / b) = b := ha.isUnit.div_div_cancel + +@[deprecated (since := "2024-11-25")] alias div_div_cancel' := div_div_cancel₀ lemma div_div_cancel_left' (ha : a ≠ 0) : a / b / a = b⁻¹ := ha.isUnit.div_div_cancel_left diff --git a/Mathlib/Algebra/Order/Field/Basic.lean b/Mathlib/Algebra/Order/Field/Basic.lean index 5832b57160d05..a8faa3f8b1daf 100644 --- a/Mathlib/Algebra/Order/Field/Basic.lean +++ b/Mathlib/Algebra/Order/Field/Basic.lean @@ -344,7 +344,7 @@ 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 := diff --git a/Mathlib/Analysis/Complex/Liouville.lean b/Mathlib/Analysis/Complex/Liouville.lean index 2d327b794e416..aee36657dab9e 100644 --- a/Mathlib/Analysis/Complex/Liouville.lean +++ b/Mathlib/Analysis/Complex/Liouville.lean @@ -93,7 +93,7 @@ theorem liouville_theorem_aux {f : ℂ → F} (hf : Differentiable ℂ f) (hb : calc ‖deriv f c‖ ≤ C / (C / ε) := norm_deriv_le_of_forall_mem_sphere_norm_le (div_pos C₀ ε₀) hf.diffContOnCl fun z _ => hC z - _ = ε := div_div_cancel' C₀.lt.ne' + _ = ε := div_div_cancel₀ C₀.lt.ne' end Complex diff --git a/Mathlib/Analysis/Complex/PhragmenLindelof.lean b/Mathlib/Analysis/Complex/PhragmenLindelof.lean index a457d182cf842..ca347d8adbcac 100644 --- a/Mathlib/Analysis/Complex/PhragmenLindelof.lean +++ b/Mathlib/Analysis/Complex/PhragmenLindelof.lean @@ -362,7 +362,7 @@ nonrec theorem quadrant_I (hd : DiffContOnCl ℂ f (Ioi 0 ×ℂ Ioi 0)) -- Porting note: failed to clear hζ ζ · -- The estimate `hB` on `f` implies the required estimate on -- `f ∘ exp` with the same `c` and `B' = max B 0`. - rw [sub_zero, div_div_cancel' Real.pi_pos.ne'] + rw [sub_zero, div_div_cancel₀ Real.pi_pos.ne'] rcases hB with ⟨c, hc, B, hO⟩ refine ⟨c, hc, max B 0, ?_⟩ rw [← comap_comap, comap_abs_atTop, comap_sup, inf_sup_right] diff --git a/Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean b/Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean index 493dc137e80c5..0dade7c5124a1 100644 --- a/Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean +++ b/Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean @@ -103,7 +103,7 @@ section symmetry /-- Reformulated functional equation with `f` and `g` interchanged. -/ lemma WeakFEPair.h_feq' (P : WeakFEPair E) (x : ℝ) (hx : 0 < x) : P.g (1 / x) = (P.ε⁻¹ * ↑(x ^ P.k)) • P.f x := by - rw [(div_div_cancel' (one_ne_zero' ℝ) ▸ P.h_feq (1 / x) (one_div_pos.mpr hx):), ← mul_smul] + rw [(div_div_cancel₀ (one_ne_zero' ℝ) ▸ P.h_feq (1 / x) (one_div_pos.mpr hx):), ← mul_smul] convert (one_smul ℂ (P.g (1 / x))).symm using 2 rw [one_div, inv_rpow hx.le, ofReal_inv] field_simp [P.hε, (rpow_pos_of_pos hx _).ne'] From 092a7d66c70504960126d34e8cc455dd042ddd91 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Wed, 27 Nov 2024 09:56:40 +0000 Subject: [PATCH 341/829] feat: coproducts in the category of Ind-objects (#19154) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Shows that - the functor `C ⥤ Ind C` preserves finite colimits, - if `C` has coproducts indexed by a finite type `α`, then `Ind C` has coproducts indexed by `α`, - if `C` has finite coproducts, then `Ind C` has small coproducts. Co-authored-by: Markus Himmel --- .../Limits/Indization/Category.lean | 83 ++++++++++++++++++- 1 file changed, 80 insertions(+), 3 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/Indization/Category.lean b/Mathlib/CategoryTheory/Limits/Indization/Category.lean index c0778f14edc72..c927eac12564d 100644 --- a/Mathlib/CategoryTheory/Limits/Indization/Category.lean +++ b/Mathlib/CategoryTheory/Limits/Indization/Category.lean @@ -3,9 +3,11 @@ Copyright (c) 2024 Markus Himmel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Markus Himmel -/ +import Mathlib.CategoryTheory.Functor.Flat +import Mathlib.CategoryTheory.Limits.Constructions.Filtered +import Mathlib.CategoryTheory.Limits.FullSubcategory import Mathlib.CategoryTheory.Limits.Indization.LocallySmall import Mathlib.CategoryTheory.Limits.Indization.FilteredColimits -import Mathlib.CategoryTheory.Limits.FullSubcategory /-! # The category of Ind-objects @@ -14,8 +16,22 @@ We define the `v`-category of Ind-objects of a category `C`, called `Ind C`, as `Ind.yoneda : C ⥤ Ind C` and `Ind.inclusion C : Ind C ⥤ Cᵒᵖ ⥤ Type v`. This file will mainly collect results about ind-objects (stated in terms of `IsIndObject`) and -reinterpret them in terms of `Ind C`. For now, we show that `Ind C` has filtered colimits and that -`Ind.inclusion C` creates them. Many other limit-colimit properties will follow. +reinterpret them in terms of `Ind C`. + +Adopting the theorem numbering of [Kashiwara2006], we show that +* `Ind C` has filtered colimits (6.1.8), +* `Ind C ⥤ Cᵒᵖ ⥤ Type v` creates filtered colimits (6.1.8), +* `C ⥤ Ind C` preserves finite colimits (6.1.6), +* if `C` has coproducts indexed by a finite type `α`, then `Ind C` has coproducts indexed by `α` + (6.1.18(ii)), +* if `C` has finite coproducts, then `Ind C` has small coproducts (6.1.18(ii)). + +More limit-colimit properties will follow. + +Note that: +* the functor `Ind C ⥤ Cᵒᵖ ⥤ Type v` does not preserve any kind of colimit in general except for + filtered colimits and +* the functor `C ⥤ Ind C` preserves finite colimits, but not infinite colimits in general. ## References * [M. Kashiwara, P. Schapira, *Categories and Sheaves*][Kashiwara2006], Chapter 6 @@ -54,6 +70,10 @@ instance : (Ind.inclusion C).Full := instance : (Ind.inclusion C).Faithful := inferInstanceAs <| ((Ind.equivalence C).functor ⋙ fullSubcategoryInclusion _).Faithful +/-- The functor `Ind C ⥤ Cᵒᵖ ⥤ Type v` is fully faithful. -/ +protected noncomputable def Ind.inclusion.fullyFaithful : (Ind.inclusion C).FullyFaithful := + .ofFullyFaithful _ + /-- The inclusion of `C` into `Ind C` induced by the Yoneda embedding. -/ protected noncomputable def Ind.yoneda : C ⥤ Ind C := FullSubcategory.lift _ CategoryTheory.yoneda isIndObject_yoneda ⋙ (Ind.equivalence C).inverse @@ -66,6 +86,10 @@ instance : (Ind.yoneda (C := C)).Faithful := inferInstanceAs <| Functor.Faithful <| FullSubcategory.lift _ CategoryTheory.yoneda isIndObject_yoneda ⋙ (Ind.equivalence C).inverse +/-- The functor `C ⥤ Ind C` is fully faithful. -/ +protected noncomputable def Ind.yoneda.fullyFaithful : (Ind.yoneda (C := C)).FullyFaithful := + .ofFullyFaithful _ + /-- The composition `C ⥤ Ind C ⥤ (Cᵒᵖ ⥤ Type v)` is just the Yoneda embedding. -/ noncomputable def Ind.yonedaCompInclusion : Ind.yoneda ⋙ Ind.inclusion C ≅ CategoryTheory.yoneda := isoWhiskerLeft (FullSubcategory.lift _ _ _) @@ -83,4 +107,57 @@ instance : HasFilteredColimits (Ind C) where HasColimitsOfShape _ _ _ := hasColimitsOfShape_of_hasColimitsOfShape_createsColimitsOfShape (Ind.inclusion C) +theorem Ind.isIndObject_inclusion_obj (X : Ind C) : IsIndObject ((Ind.inclusion C).obj X) := + X.2 + +/-- Pick a presentation of an ind-object `X` using choice. -/ +noncomputable def Ind.presentation (X : Ind C) : IndObjectPresentation ((Ind.inclusion C).obj X) := + X.isIndObject_inclusion_obj.presentation + +/-- An ind-object `X` is the colimit (in `Ind C`!) of the filtered diagram presenting it. -/ +noncomputable def Ind.colimitPresentationCompYoneda (X : Ind C) : + colimit (X.presentation.F ⋙ Ind.yoneda) ≅ X := + Ind.inclusion.fullyFaithful.isoEquiv.symm <| calc + (Ind.inclusion C).obj (colimit (X.presentation.F ⋙ Ind.yoneda)) + ≅ colimit (X.presentation.F ⋙ Ind.yoneda ⋙ Ind.inclusion C) := preservesColimitIso _ _ + _ ≅ colimit (X.presentation.F ⋙ yoneda) := + HasColimit.isoOfNatIso (isoWhiskerLeft X.presentation.F Ind.yonedaCompInclusion) + _ ≅ (Ind.inclusion C).obj X := + IsColimit.coconePointUniqueUpToIso (colimit.isColimit _) X.presentation.isColimit + +instance : RepresentablyCoflat (Ind.yoneda (C := C)) := by + refine ⟨fun X => ?_⟩ + suffices IsFiltered (CostructuredArrow yoneda ((Ind.inclusion C).obj X)) from + IsFiltered.of_equivalence + ((CostructuredArrow.post Ind.yoneda (Ind.inclusion C) X).asEquivalence.trans + (CostructuredArrow.mapNatIso Ind.yonedaCompInclusion)).symm + exact ((isIndObject_iff _).1 (Ind.isIndObject_inclusion_obj X)).1 + +noncomputable instance : PreservesFiniteColimits (Ind.yoneda (C := C)) := + preservesFiniteColimits_of_coflat _ + +instance {α : Type v} [Finite α] [HasColimitsOfShape (Discrete α) C] : + HasColimitsOfShape (Discrete α) (Ind C) := by + refine ⟨fun F => ?_⟩ + let I : α → Type v := fun s => (F.obj ⟨s⟩).presentation.I + let G : ∀ s, I s ⥤ C := fun s => (F.obj ⟨s⟩).presentation.F + let iso : Discrete.functor (fun s => Pi.eval I s ⋙ G s) ⋙ + (whiskeringRight _ _ _).obj Ind.yoneda ⋙ colim ≅ F := by + refine Discrete.natIso (fun s => ?_) + refine (Functor.Final.colimitIso (Pi.eval I s.as) (G s.as ⋙ Ind.yoneda)) ≪≫ ?_ + exact Ind.colimitPresentationCompYoneda _ + -- The actual proof happens during typeclass resolution in the following line, which deduces + -- ``` + -- HasColimit Discrete.functor (fun s => Pi.eval I s ⋙ G s) ⋙ + -- (whiskeringRight _ _ _).obj Ind.yoneda ⋙ colim + -- ``` + -- from the fact that finite limits commute with filtered colimits and from the fact that + -- `Ind.yoneda` preserves finite colimits. + apply hasColimitOfIso iso.symm + +instance [HasFiniteCoproducts C] : HasCoproducts.{v} (Ind C) := + have : HasFiniteCoproducts (Ind C) := + ⟨fun _ => hasColimitsOfShape_of_equivalence (Discrete.equivalence Equiv.ulift)⟩ + hasCoproducts_of_finite_and_filtered + end CategoryTheory From 134190977464360ecc5de1cd85ffb25b1172db7d Mon Sep 17 00:00:00 2001 From: grhkm21 Date: Wed, 27 Nov 2024 09:56:41 +0000 Subject: [PATCH 342/829] feat: prove `zmodRepr` is unique (#19391) We prove `PadicInt.exists_unique_mem_range` which is a strengthening of `exists_mem_range`, proving uniqueness of the natural number. I need this to simplify `toZMod(1+ap)` to `1`. --- Mathlib/NumberTheory/Padics/RingHoms.lean | 14 ++++++++++++-- 1 file changed, 12 insertions(+), 2 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/RingHoms.lean b/Mathlib/NumberTheory/Padics/RingHoms.lean index b80610a74c52a..2fed1c971dbfe 100644 --- a/Mathlib/NumberTheory/Padics/RingHoms.lean +++ b/Mathlib/NumberTheory/Padics/RingHoms.lean @@ -178,14 +178,24 @@ theorem exists_mem_range : ∃ n : ℕ, n < p ∧ x - n ∈ maximalIdeal ℤ_[p] apply max_lt hr simpa using hn +theorem exists_unique_mem_range : ∃! n : ℕ, n < p ∧ x - n ∈ maximalIdeal ℤ_[p] := by + obtain ⟨n, hn₁, hn₂⟩ := exists_mem_range x + use n, ⟨hn₁, hn₂⟩, fun m ⟨hm₁, hm₂⟩ ↦ ?_ + have := (zmod_congr_of_sub_mem_max_ideal x n m hn₂ hm₂).symm + rwa [ZMod.natCast_eq_natCast_iff, ModEq, mod_eq_of_lt hn₁, mod_eq_of_lt hm₁] at this + /-- `zmod_repr x` is the unique natural number smaller than `p` satisfying `‖(x - zmod_repr x : ℤ_[p])‖ < 1`. -/ def zmodRepr : ℕ := - Classical.choose (exists_mem_range x) + Classical.choose (exists_unique_mem_range x).exists theorem zmodRepr_spec : zmodRepr x < p ∧ x - zmodRepr x ∈ maximalIdeal ℤ_[p] := - Classical.choose_spec (exists_mem_range x) + Classical.choose_spec (exists_unique_mem_range x).exists + +theorem zmodRepr_unique (y : ℕ) (hy₁ : y < p) (hy₂ : x - y ∈ maximalIdeal ℤ_[p]) : y = zmodRepr x := + have h := (Classical.choose_spec (exists_unique_mem_range x)).right + (h y ⟨hy₁, hy₂⟩).trans (h (zmodRepr x) (zmodRepr_spec x)).symm theorem zmodRepr_lt_p : zmodRepr x < p := (zmodRepr_spec _).1 From 8674145a5b085f53aafb08bca0fd7eae8e9a1d9d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 27 Nov 2024 09:56:42 +0000 Subject: [PATCH 343/829] chore(NumberTheory/Fermat): fix theorem names (#19469) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From LeanCamCombi Co-authored-by: Yves Jäckle --- Mathlib/NumberTheory/Fermat.lean | 49 ++++++++++++++++++++------------ 1 file changed, 31 insertions(+), 18 deletions(-) diff --git a/Mathlib/NumberTheory/Fermat.lean b/Mathlib/NumberTheory/Fermat.lean index bdb8dd1117803..c2aca1d62719d 100644 --- a/Mathlib/NumberTheory/Fermat.lean +++ b/Mathlib/NumberTheory/Fermat.lean @@ -20,6 +20,8 @@ for all natural numbers `n`. form `k * 2 ^ (n + 2) + 1`. -/ +open Function + namespace Nat open Finset Nat ZMod @@ -32,20 +34,25 @@ def fermatNumber (n : ℕ) : ℕ := 2 ^ (2 ^ n) + 1 @[simp] theorem fermatNumber_one : fermatNumber 1 = 5 := rfl @[simp] theorem fermatNumber_two : fermatNumber 2 = 17 := rfl -theorem strictMono_fermatNumber : StrictMono fermatNumber := by +theorem fermatNumber_strictMono : StrictMono fermatNumber := by intro m n simp only [fermatNumber, add_lt_add_iff_right, Nat.pow_lt_pow_iff_right (one_lt_two : 1 < 2), imp_self] -theorem two_lt_fermatNumber (n : ℕ) : 2 < fermatNumber n := by - cases n - · simp - · exact lt_of_succ_lt <| strictMono_fermatNumber <| zero_lt_succ _ +@[deprecated (since := "2024-11-25")] alias strictMono_fermatNumber := fermatNumber_strictMono + +lemma fermatNumber_mono : Monotone fermatNumber := fermatNumber_strictMono.monotone +lemma fermatNumber_injective : Injective fermatNumber := fermatNumber_strictMono.injective + +lemma three_le_fermatNumber (n : ℕ) : 3 ≤ fermatNumber n := fermatNumber_mono n.zero_le +lemma two_lt_fermatNumber (n : ℕ) : 2 < fermatNumber n := three_le_fermatNumber _ + +lemma fermatNumber_ne_one (n : ℕ) : fermatNumber n ≠ 1 := by have := three_le_fermatNumber n; omega theorem odd_fermatNumber (n : ℕ) : Odd (fermatNumber n) := (even_pow.mpr ⟨even_two, (pow_pos two_pos n).ne'⟩).add_one -theorem fermatNumber_product (n : ℕ) : ∏ k in range n, fermatNumber k = fermatNumber n - 2 := by +theorem prod_fermatNumber (n : ℕ) : ∏ k ∈ range n, fermatNumber k = fermatNumber n - 2 := by induction' n with n hn · rfl rw [prod_range_succ, hn, fermatNumber, fermatNumber, mul_comm, @@ -53,9 +60,11 @@ theorem fermatNumber_product (n : ℕ) : ∏ k in range n, fermatNumber k = ferm ring_nf omega +@[deprecated (since := "2024-11-25")] alias fermatNumber_product := prod_fermatNumber + theorem fermatNumber_eq_prod_add_two (n : ℕ) : - fermatNumber n = (∏ k in range n, fermatNumber k) + 2 := by - rw [fermatNumber_product, Nat.sub_add_cancel] + fermatNumber n = ∏ k ∈ range n, fermatNumber k + 2 := by + rw [prod_fermatNumber, Nat.sub_add_cancel] exact le_of_lt <| two_lt_fermatNumber _ theorem fermatNumber_succ (n : ℕ) : fermatNumber (n + 1) = (fermatNumber n - 1) ^ 2 + 1 := by @@ -93,16 +102,20 @@ open Finset From a letter to Euler, see page 37 in [juskevic2022]. -/ -theorem coprime_fermatNumber_fermatNumber {k n : ℕ} (h : k ≠ n) : - Coprime (fermatNumber n) (fermatNumber k) := by - wlog hkn : k < n - · simpa only [coprime_comm] using this h.symm (by omega) - let m := (fermatNumber n).gcd (fermatNumber k) - have h_n : m ∣ fermatNumber n := (fermatNumber n).gcd_dvd_left (fermatNumber k) - have h_m : m ∣ 2 := (Nat.dvd_add_right <| (gcd_dvd_right _ _).trans <| dvd_prod_of_mem _ - <| mem_range.mpr hkn).mp <| fermatNumber_eq_prod_add_two _ ▸ h_n - refine ((dvd_prime prime_two).mp h_m).elim id (fun h_two ↦ ?_) - exact ((odd_fermatNumber _).not_two_dvd_nat (h_two ▸ h_n)).elim +theorem coprime_fermatNumber_fermatNumber {m n : ℕ} (hmn : m ≠ n) : + Coprime (fermatNumber m) (fermatNumber n) := by + wlog hmn' : m < n + · simpa only [coprime_comm] using this hmn.symm (by omega) + let d := (fermatNumber m).gcd (fermatNumber n) + have h_n : d ∣ fermatNumber n := gcd_dvd_right .. + have h_m : d ∣ 2 := (Nat.dvd_add_right <| (gcd_dvd_left _ _).trans <| dvd_prod_of_mem _ + <| mem_range.mpr hmn').mp <| fermatNumber_eq_prod_add_two _ ▸ h_n + refine ((dvd_prime prime_two).mp h_m).resolve_right fun h_two ↦ ?_ + exact (odd_fermatNumber _).not_two_dvd_nat (h_two ▸ h_n) + +lemma pairwise_coprime_fermatNumber : + Pairwise fun m n ↦ Coprime (fermatNumber m) (fermatNumber n) := + fun _m _n ↦ coprime_fermatNumber_fermatNumber open ZMod From 00200113ae7e6500fe14d089684b57daef691d0e Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 27 Nov 2024 09:56:44 +0000 Subject: [PATCH 344/829] chore(GroupTheory/FreeGroup): split off `Reduce.lean` (#19501) --- Mathlib.lean | 1 + Mathlib/GroupTheory/FreeGroup/Basic.lean | 326 +------------------- Mathlib/GroupTheory/FreeGroup/Reduce.lean | 351 ++++++++++++++++++++++ Mathlib/SetTheory/Cardinal/Free.lean | 1 + 4 files changed, 354 insertions(+), 325 deletions(-) create mode 100644 Mathlib/GroupTheory/FreeGroup/Reduce.lean diff --git a/Mathlib.lean b/Mathlib.lean index 21358bdd55951..ff6ae6de6f371 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3115,6 +3115,7 @@ import Mathlib.GroupTheory.FreeAbelianGroupFinsupp import Mathlib.GroupTheory.FreeGroup.Basic import Mathlib.GroupTheory.FreeGroup.IsFreeGroup import Mathlib.GroupTheory.FreeGroup.NielsenSchreier +import Mathlib.GroupTheory.FreeGroup.Reduce import Mathlib.GroupTheory.GroupAction.Basic import Mathlib.GroupTheory.GroupAction.Blocks import Mathlib.GroupTheory.GroupAction.CardCommute diff --git a/Mathlib/GroupTheory/FreeGroup/Basic.lean b/Mathlib/GroupTheory/FreeGroup/Basic.lean index 2c47200b1ac4d..a2174bda20dfc 100644 --- a/Mathlib/GroupTheory/FreeGroup/Basic.lean +++ b/Mathlib/GroupTheory/FreeGroup/Basic.lean @@ -3,9 +3,8 @@ Copyright (c) 2018 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ +import Mathlib.Algebra.BigOperators.Group.List import Mathlib.Algebra.Group.Subgroup.Ker -import Mathlib.Data.Fintype.Basic -import Mathlib.Data.List.Sublists /-! # Free groups @@ -895,327 +894,4 @@ instance : LawfulMonad FreeGroup.{u} := LawfulMonad.mk' end Category -section Reduce - -variable [DecidableEq α] - -/-- The maximal reduction of a word. It is computable -iff `α` has decidable equality. -/ -@[to_additive "The maximal reduction of a word. It is computable iff `α` has decidable equality."] -def reduce : (L : List (α × Bool)) -> List (α × Bool) := - List.rec [] fun hd1 _tl1 ih => - List.casesOn ih [hd1] fun hd2 tl2 => - if hd1.1 = hd2.1 ∧ hd1.2 = not hd2.2 then tl2 else hd1 :: hd2 :: tl2 - -@[to_additive (attr := simp)] lemma reduce_nil : reduce ([] : List (α × Bool)) = [] := rfl -@[to_additive] lemma reduce_singleton (s : α × Bool) : reduce [s] = [s] := rfl - -@[to_additive (attr := simp)] -theorem reduce.cons (x) : - reduce (x :: L) = - List.casesOn (reduce L) [x] fun hd tl => - if x.1 = hd.1 ∧ x.2 = not hd.2 then tl else x :: hd :: tl := - rfl - -@[to_additive (attr := simp)] -theorem reduce_replicate (n : ℕ) (x : α × Bool) : - reduce (.replicate n x) = .replicate n x := by - induction n with - | zero => simp [reduce] - | succ n ih => - rw [List.replicate_succ, reduce.cons, ih] - cases n with - | zero => simp - | succ n => simp [List.replicate_succ] - -/-- The first theorem that characterises the function `reduce`: a word reduces to its maximal - reduction. -/ -@[to_additive "The first theorem that characterises the function `reduce`: a word reduces to its - maximal reduction."] -theorem reduce.red : Red L (reduce L) := by - induction L with - | nil => constructor - | cons hd1 tl1 ih => - dsimp - revert ih - generalize htl : reduce tl1 = TL - intro ih - cases TL with - | nil => exact Red.cons_cons ih - | cons hd2 tl2 => - dsimp only - split_ifs with h - · cases hd1 - cases hd2 - cases h - dsimp at * - subst_vars - apply Red.trans (Red.cons_cons ih) - exact Red.Step.cons_not_rev.to_red - · exact Red.cons_cons ih - -@[to_additive] -theorem reduce.not {p : Prop} : - ∀ {L₁ L₂ L₃ : List (α × Bool)} {x b}, reduce L₁ = L₂ ++ (x, b) :: (x, !b) :: L₃ → p - | [], L2, L3, _, _ => fun h => by cases L2 <;> injections - | (x, b) :: L1, L2, L3, x', b' => by - dsimp - cases r : reduce L1 with - | nil => - dsimp; intro h - exfalso - have := congr_arg List.length h - simp? [List.length] at this says - simp only [List.length, zero_add, List.length_append] at this - rw [add_comm, add_assoc, add_assoc, add_comm, <-add_assoc] at this - omega - | cons hd tail => - cases' hd with y c - dsimp only - split_ifs with h <;> intro H - · rw [H] at r - exact @reduce.not _ L1 ((y, c) :: L2) L3 x' b' r - rcases L2 with (_ | ⟨a, L2⟩) - · injections; subst_vars - simp at h - · refine @reduce.not _ L1 L2 L3 x' b' ?_ - injection H with _ H - rw [r, H]; rfl - -/-- The second theorem that characterises the function `reduce`: the maximal reduction of a word -only reduces to itself. -/ -@[to_additive "The second theorem that characterises the function `reduce`: the maximal reduction of - a word only reduces to itself."] -theorem reduce.min (H : Red (reduce L₁) L₂) : reduce L₁ = L₂ := by - induction' H with L1 L' L2 H1 H2 ih - · rfl - · cases' H1 with L4 L5 x b - exact reduce.not H2 - -/-- `reduce` is idempotent, i.e. the maximal reduction of the maximal reduction of a word is the - maximal reduction of the word. -/ -@[to_additive (attr := simp) "`reduce` is idempotent, i.e. the maximal reduction of the maximal - reduction of a word is the maximal reduction of the word."] -theorem reduce.idem : reduce (reduce L) = reduce L := - Eq.symm <| reduce.min reduce.red - -@[to_additive] -theorem reduce.Step.eq (H : Red.Step L₁ L₂) : reduce L₁ = reduce L₂ := - let ⟨_L₃, HR13, HR23⟩ := Red.church_rosser reduce.red (reduce.red.head H) - (reduce.min HR13).trans (reduce.min HR23).symm - -/-- If a word reduces to another word, then they have a common maximal reduction. -/ -@[to_additive "If a word reduces to another word, then they have a common maximal reduction."] -theorem reduce.eq_of_red (H : Red L₁ L₂) : reduce L₁ = reduce L₂ := - let ⟨_L₃, HR13, HR23⟩ := Red.church_rosser reduce.red (Red.trans H reduce.red) - (reduce.min HR13).trans (reduce.min HR23).symm - -alias red.reduce_eq := reduce.eq_of_red - -alias freeAddGroup.red.reduce_eq := FreeAddGroup.reduce.eq_of_red - -@[to_additive] -theorem Red.reduce_right (h : Red L₁ L₂) : Red L₁ (reduce L₂) := - reduce.eq_of_red h ▸ reduce.red - -@[to_additive] -theorem Red.reduce_left (h : Red L₁ L₂) : Red L₂ (reduce L₁) := - (reduce.eq_of_red h).symm ▸ reduce.red - -/-- If two words correspond to the same element in the free group, then they -have a common maximal reduction. This is the proof that the function that sends -an element of the free group to its maximal reduction is well-defined. -/ -@[to_additive "If two words correspond to the same element in the additive free group, then they - have a common maximal reduction. This is the proof that the function that sends an element of the - free group to its maximal reduction is well-defined."] -theorem reduce.sound (H : mk L₁ = mk L₂) : reduce L₁ = reduce L₂ := - let ⟨_L₃, H13, H23⟩ := Red.exact.1 H - (reduce.eq_of_red H13).trans (reduce.eq_of_red H23).symm - -/-- If two words have a common maximal reduction, then they correspond to the same element in the - free group. -/ -@[to_additive "If two words have a common maximal reduction, then they correspond to the same - element in the additive free group."] -theorem reduce.exact (H : reduce L₁ = reduce L₂) : mk L₁ = mk L₂ := - Red.exact.2 ⟨reduce L₂, H ▸ reduce.red, reduce.red⟩ - -/-- A word and its maximal reduction correspond to the same element of the free group. -/ -@[to_additive "A word and its maximal reduction correspond to the same element of the additive free - group."] -theorem reduce.self : mk (reduce L) = mk L := - reduce.exact reduce.idem - -/-- If words `w₁ w₂` are such that `w₁` reduces to `w₂`, then `w₂` reduces to the maximal reduction - of `w₁`. -/ -@[to_additive "If words `w₁ w₂` are such that `w₁` reduces to `w₂`, then `w₂` reduces to the maximal - reduction of `w₁`."] -theorem reduce.rev (H : Red L₁ L₂) : Red L₂ (reduce L₁) := - (reduce.eq_of_red H).symm ▸ reduce.red - -/-- The function that sends an element of the free group to its maximal reduction. -/ -@[to_additive "The function that sends an element of the additive free group to its maximal - reduction."] -def toWord : FreeGroup α → List (α × Bool) := - Quot.lift reduce fun _L₁ _L₂ H => reduce.Step.eq H - -@[to_additive] -theorem mk_toWord : ∀ {x : FreeGroup α}, mk (toWord x) = x := by rintro ⟨L⟩; exact reduce.self - -@[to_additive] -theorem toWord_injective : Function.Injective (toWord : FreeGroup α → List (α × Bool)) := by - rintro ⟨L₁⟩ ⟨L₂⟩; exact reduce.exact - -@[to_additive (attr := simp)] -theorem toWord_inj {x y : FreeGroup α} : toWord x = toWord y ↔ x = y := - toWord_injective.eq_iff - -@[to_additive (attr := simp)] -theorem toWord_mk : (mk L₁).toWord = reduce L₁ := - rfl - -@[to_additive (attr := simp)] -theorem toWord_of (a : α) : (of a).toWord = [(a, true)] := - rfl - -@[to_additive (attr := simp)] -theorem reduce_toWord : ∀ x : FreeGroup α, reduce (toWord x) = toWord x := by - rintro ⟨L⟩ - exact reduce.idem - -@[to_additive (attr := simp)] -theorem toWord_one : (1 : FreeGroup α).toWord = [] := - rfl - -@[to_additive (attr := simp)] -theorem toWord_of_pow (a : α) (n : ℕ) : (of a ^ n).toWord = List.replicate n (a, true) := by - rw [of, pow_mk, List.flatten_replicate_singleton, toWord] - exact reduce_replicate _ _ - -@[to_additive (attr := simp)] -theorem toWord_eq_nil_iff {x : FreeGroup α} : x.toWord = [] ↔ x = 1 := - toWord_injective.eq_iff' toWord_one - -@[to_additive] -theorem reduce_invRev {w : List (α × Bool)} : reduce (invRev w) = invRev (reduce w) := by - apply reduce.min - rw [← red_invRev_iff, invRev_invRev] - apply Red.reduce_left - have : Red (invRev (invRev w)) (invRev (reduce (invRev w))) := reduce.red.invRev - rwa [invRev_invRev] at this - -@[to_additive (attr := simp)] -theorem toWord_inv (x : FreeGroup α) : x⁻¹.toWord = invRev x.toWord := by - rcases x with ⟨L⟩ - rw [quot_mk_eq_mk, inv_mk, toWord_mk, toWord_mk, reduce_invRev] - -@[to_additive] -lemma toWord_mul_sublist (x y : FreeGroup α) : (x * y).toWord <+ x.toWord ++ y.toWord := by - refine Red.sublist ?_ - have : x * y = FreeGroup.mk (x.toWord ++ y.toWord) := by - rw [← FreeGroup.mul_mk, FreeGroup.mk_toWord, FreeGroup.mk_toWord] - rw [this] - exact FreeGroup.reduce.red - -/-- **Constructive Church-Rosser theorem** (compare `church_rosser`). -/ -@[to_additive "**Constructive Church-Rosser theorem** (compare `church_rosser`)."] -def reduce.churchRosser (H12 : Red L₁ L₂) (H13 : Red L₁ L₃) : { L₄ // Red L₂ L₄ ∧ Red L₃ L₄ } := - ⟨reduce L₁, reduce.rev H12, reduce.rev H13⟩ - -@[to_additive] -instance : DecidableEq (FreeGroup α) := - toWord_injective.decidableEq - --- TODO @[to_additive] doesn't succeed, possibly due to a bug -instance Red.decidableRel : DecidableRel (@Red α) - | [], [] => isTrue Red.refl - | [], _hd2 :: _tl2 => isFalse fun H => List.noConfusion (Red.nil_iff.1 H) - | (x, b) :: tl, [] => - match Red.decidableRel tl [(x, not b)] with - | isTrue H => isTrue <| Red.trans (Red.cons_cons H) <| (@Red.Step.not _ [] [] _ _).to_red - | isFalse H => isFalse fun H2 => H <| Red.cons_nil_iff_singleton.1 H2 - | (x1, b1) :: tl1, (x2, b2) :: tl2 => - if h : (x1, b1) = (x2, b2) then - match Red.decidableRel tl1 tl2 with - | isTrue H => isTrue <| h ▸ Red.cons_cons H - | isFalse H => isFalse fun H2 => H <| (Red.cons_cons_iff _).1 <| h.symm ▸ H2 - else - match Red.decidableRel tl1 ((x1, ! b1) :: (x2, b2) :: tl2) with - | isTrue H => isTrue <| (Red.cons_cons H).tail Red.Step.cons_not - | isFalse H => isFalse fun H2 => H <| Red.inv_of_red_of_ne h H2 - -/-- A list containing every word that `w₁` reduces to. -/ -def Red.enum (L₁ : List (α × Bool)) : List (List (α × Bool)) := - List.filter (Red L₁) (List.sublists L₁) - -theorem Red.enum.sound (H : L₂ ∈ List.filter (Red L₁) (List.sublists L₁)) : Red L₁ L₂ := - of_decide_eq_true (@List.of_mem_filter _ _ L₂ _ H) - -theorem Red.enum.complete (H : Red L₁ L₂) : L₂ ∈ Red.enum L₁ := - List.mem_filter_of_mem (List.mem_sublists.2 <| Red.sublist H) (decide_eq_true H) - -instance : Fintype { L₂ // Red L₁ L₂ } := - Fintype.subtype (List.toFinset <| Red.enum L₁) fun _L₂ => - ⟨fun H => Red.enum.sound <| List.mem_toFinset.1 H, fun H => - List.mem_toFinset.2 <| Red.enum.complete H⟩ - -end Reduce - -@[to_additive (attr := simp)] -theorem one_ne_of (a : α) : 1 ≠ of a := - letI := Classical.decEq α; ne_of_apply_ne toWord <| by simp - -@[to_additive (attr := simp)] -theorem of_ne_one (a : α) : of a ≠ 1 := one_ne_of _ |>.symm - -@[to_additive] -instance [Nonempty α] : Nontrivial (FreeGroup α) where - exists_pair_ne := let ⟨x⟩ := ‹Nonempty α›; ⟨1, of x, one_ne_of x⟩ - -section Metric - -variable [DecidableEq α] - -/-- The length of reduced words provides a norm on a free group. -/ -@[to_additive "The length of reduced words provides a norm on an additive free group."] -def norm (x : FreeGroup α) : ℕ := - x.toWord.length - -@[to_additive (attr := simp)] -theorem norm_inv_eq {x : FreeGroup α} : norm x⁻¹ = norm x := by - simp only [norm, toWord_inv, invRev_length] - -@[to_additive (attr := simp)] -theorem norm_eq_zero {x : FreeGroup α} : norm x = 0 ↔ x = 1 := by - simp only [norm, List.length_eq_zero, toWord_eq_nil_iff] - -@[to_additive (attr := simp)] -theorem norm_one : norm (1 : FreeGroup α) = 0 := - rfl - -@[to_additive (attr := simp)] -theorem norm_of (a : α) : norm (of a) = 1 := - rfl - -@[to_additive] -theorem norm_mk_le : norm (mk L₁) ≤ L₁.length := - reduce.red.length_le - -@[to_additive] -theorem norm_mul_le (x y : FreeGroup α) : norm (x * y) ≤ norm x + norm y := - calc - norm (x * y) = norm (mk (x.toWord ++ y.toWord)) := by rw [← mul_mk, mk_toWord, mk_toWord] - _ ≤ (x.toWord ++ y.toWord).length := norm_mk_le - _ = norm x + norm y := List.length_append _ _ - -@[to_additive (attr := simp)] -theorem norm_of_pow (a : α) (n : ℕ) : norm (of a ^ n) = n := by - rw [norm, toWord_of_pow, List.length_replicate] - -@[to_additive] -theorem norm_surjective [Nonempty α] : Function.Surjective (norm (α := α)) := by - let ⟨a⟩ := ‹Nonempty α› - exact Function.RightInverse.surjective <| norm_of_pow a - -end Metric - end FreeGroup diff --git a/Mathlib/GroupTheory/FreeGroup/Reduce.lean b/Mathlib/GroupTheory/FreeGroup/Reduce.lean new file mode 100644 index 0000000000000..3017c17281dd2 --- /dev/null +++ b/Mathlib/GroupTheory/FreeGroup/Reduce.lean @@ -0,0 +1,351 @@ +/- +Copyright (c) 2018 Kenny Lau. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kenny Lau +-/ +import Mathlib.GroupTheory.FreeGroup.Basic +import Mathlib.Data.Fintype.Basic +import Mathlib.Data.List.Sublists + +/-! +# The maximal reduction of a word in a free group + +## Main declarations + +* `FreeGroup.reduce`: the maximal reduction of a word in a free group +* `FreeGroup.norm`: the length of the maximal reduction of a word in a free group + +-/ + + +namespace FreeGroup + +variable {α : Type*} +variable {L L₁ L₂ L₃ L₄ : List (α × Bool)} + +section Reduce + +variable [DecidableEq α] + +/-- The maximal reduction of a word. It is computable +iff `α` has decidable equality. -/ +@[to_additive "The maximal reduction of a word. It is computable iff `α` has decidable equality."] +def reduce : (L : List (α × Bool)) -> List (α × Bool) := + List.rec [] fun hd1 _tl1 ih => + List.casesOn ih [hd1] fun hd2 tl2 => + if hd1.1 = hd2.1 ∧ hd1.2 = not hd2.2 then tl2 else hd1 :: hd2 :: tl2 + +@[to_additive (attr := simp)] lemma reduce_nil : reduce ([] : List (α × Bool)) = [] := rfl +@[to_additive] lemma reduce_singleton (s : α × Bool) : reduce [s] = [s] := rfl + +@[to_additive (attr := simp)] +theorem reduce.cons (x) : + reduce (x :: L) = + List.casesOn (reduce L) [x] fun hd tl => + if x.1 = hd.1 ∧ x.2 = not hd.2 then tl else x :: hd :: tl := + rfl + +@[to_additive (attr := simp)] +theorem reduce_replicate (n : ℕ) (x : α × Bool) : + reduce (.replicate n x) = .replicate n x := by + induction n with + | zero => simp [reduce] + | succ n ih => + rw [List.replicate_succ, reduce.cons, ih] + cases n with + | zero => simp + | succ n => simp [List.replicate_succ] + +/-- The first theorem that characterises the function `reduce`: a word reduces to its maximal + reduction. -/ +@[to_additive "The first theorem that characterises the function `reduce`: a word reduces to its + maximal reduction."] +theorem reduce.red : Red L (reduce L) := by + induction L with + | nil => constructor + | cons hd1 tl1 ih => + dsimp + revert ih + generalize htl : reduce tl1 = TL + intro ih + cases TL with + | nil => exact Red.cons_cons ih + | cons hd2 tl2 => + dsimp only + split_ifs with h + · cases hd1 + cases hd2 + cases h + dsimp at * + subst_vars + apply Red.trans (Red.cons_cons ih) + exact Red.Step.cons_not_rev.to_red + · exact Red.cons_cons ih + +@[to_additive] +theorem reduce.not {p : Prop} : + ∀ {L₁ L₂ L₃ : List (α × Bool)} {x b}, reduce L₁ = L₂ ++ (x, b) :: (x, !b) :: L₃ → p + | [], L2, L3, _, _ => fun h => by cases L2 <;> injections + | (x, b) :: L1, L2, L3, x', b' => by + dsimp + cases r : reduce L1 with + | nil => + dsimp; intro h + exfalso + have := congr_arg List.length h + simp? [List.length] at this says + simp only [List.length, zero_add, List.length_append] at this + rw [add_comm, add_assoc, add_assoc, add_comm, <-add_assoc] at this + omega + | cons hd tail => + cases' hd with y c + dsimp only + split_ifs with h <;> intro H + · rw [H] at r + exact @reduce.not _ L1 ((y, c) :: L2) L3 x' b' r + rcases L2 with (_ | ⟨a, L2⟩) + · injections; subst_vars + simp at h + · refine @reduce.not _ L1 L2 L3 x' b' ?_ + injection H with _ H + rw [r, H]; rfl + +/-- The second theorem that characterises the function `reduce`: the maximal reduction of a word +only reduces to itself. -/ +@[to_additive "The second theorem that characterises the function `reduce`: the maximal reduction of + a word only reduces to itself."] +theorem reduce.min (H : Red (reduce L₁) L₂) : reduce L₁ = L₂ := by + induction' H with L1 L' L2 H1 H2 ih + · rfl + · cases' H1 with L4 L5 x b + exact reduce.not H2 + +/-- `reduce` is idempotent, i.e. the maximal reduction of the maximal reduction of a word is the + maximal reduction of the word. -/ +@[to_additive (attr := simp) "`reduce` is idempotent, i.e. the maximal reduction of the maximal + reduction of a word is the maximal reduction of the word."] +theorem reduce.idem : reduce (reduce L) = reduce L := + Eq.symm <| reduce.min reduce.red + +@[to_additive] +theorem reduce.Step.eq (H : Red.Step L₁ L₂) : reduce L₁ = reduce L₂ := + let ⟨_L₃, HR13, HR23⟩ := Red.church_rosser reduce.red (reduce.red.head H) + (reduce.min HR13).trans (reduce.min HR23).symm + +/-- If a word reduces to another word, then they have a common maximal reduction. -/ +@[to_additive "If a word reduces to another word, then they have a common maximal reduction."] +theorem reduce.eq_of_red (H : Red L₁ L₂) : reduce L₁ = reduce L₂ := + let ⟨_L₃, HR13, HR23⟩ := Red.church_rosser reduce.red (Red.trans H reduce.red) + (reduce.min HR13).trans (reduce.min HR23).symm + +alias red.reduce_eq := reduce.eq_of_red + +alias freeAddGroup.red.reduce_eq := FreeAddGroup.reduce.eq_of_red + +@[to_additive] +theorem Red.reduce_right (h : Red L₁ L₂) : Red L₁ (reduce L₂) := + reduce.eq_of_red h ▸ reduce.red + +@[to_additive] +theorem Red.reduce_left (h : Red L₁ L₂) : Red L₂ (reduce L₁) := + (reduce.eq_of_red h).symm ▸ reduce.red + +/-- If two words correspond to the same element in the free group, then they +have a common maximal reduction. This is the proof that the function that sends +an element of the free group to its maximal reduction is well-defined. -/ +@[to_additive "If two words correspond to the same element in the additive free group, then they + have a common maximal reduction. This is the proof that the function that sends an element of the + free group to its maximal reduction is well-defined."] +theorem reduce.sound (H : mk L₁ = mk L₂) : reduce L₁ = reduce L₂ := + let ⟨_L₃, H13, H23⟩ := Red.exact.1 H + (reduce.eq_of_red H13).trans (reduce.eq_of_red H23).symm + +/-- If two words have a common maximal reduction, then they correspond to the same element in the + free group. -/ +@[to_additive "If two words have a common maximal reduction, then they correspond to the same + element in the additive free group."] +theorem reduce.exact (H : reduce L₁ = reduce L₂) : mk L₁ = mk L₂ := + Red.exact.2 ⟨reduce L₂, H ▸ reduce.red, reduce.red⟩ + +/-- A word and its maximal reduction correspond to the same element of the free group. -/ +@[to_additive "A word and its maximal reduction correspond to the same element of the additive free + group."] +theorem reduce.self : mk (reduce L) = mk L := + reduce.exact reduce.idem + +/-- If words `w₁ w₂` are such that `w₁` reduces to `w₂`, then `w₂` reduces to the maximal reduction + of `w₁`. -/ +@[to_additive "If words `w₁ w₂` are such that `w₁` reduces to `w₂`, then `w₂` reduces to the maximal + reduction of `w₁`."] +theorem reduce.rev (H : Red L₁ L₂) : Red L₂ (reduce L₁) := + (reduce.eq_of_red H).symm ▸ reduce.red + +/-- The function that sends an element of the free group to its maximal reduction. -/ +@[to_additive "The function that sends an element of the additive free group to its maximal + reduction."] +def toWord : FreeGroup α → List (α × Bool) := + Quot.lift reduce fun _L₁ _L₂ H => reduce.Step.eq H + +@[to_additive] +theorem mk_toWord : ∀ {x : FreeGroup α}, mk (toWord x) = x := by rintro ⟨L⟩; exact reduce.self + +@[to_additive] +theorem toWord_injective : Function.Injective (toWord : FreeGroup α → List (α × Bool)) := by + rintro ⟨L₁⟩ ⟨L₂⟩; exact reduce.exact + +@[to_additive (attr := simp)] +theorem toWord_inj {x y : FreeGroup α} : toWord x = toWord y ↔ x = y := + toWord_injective.eq_iff + +@[to_additive (attr := simp)] +theorem toWord_mk : (mk L₁).toWord = reduce L₁ := + rfl + +@[to_additive (attr := simp)] +theorem toWord_of (a : α) : (of a).toWord = [(a, true)] := + rfl + +@[to_additive (attr := simp)] +theorem reduce_toWord : ∀ x : FreeGroup α, reduce (toWord x) = toWord x := by + rintro ⟨L⟩ + exact reduce.idem + +@[to_additive (attr := simp)] +theorem toWord_one : (1 : FreeGroup α).toWord = [] := + rfl + +@[to_additive (attr := simp)] +theorem toWord_of_pow (a : α) (n : ℕ) : (of a ^ n).toWord = List.replicate n (a, true) := by + rw [of, pow_mk, List.flatten_replicate_singleton, toWord] + exact reduce_replicate _ _ + +@[to_additive (attr := simp)] +theorem toWord_eq_nil_iff {x : FreeGroup α} : x.toWord = [] ↔ x = 1 := + toWord_injective.eq_iff' toWord_one + +@[to_additive] +theorem reduce_invRev {w : List (α × Bool)} : reduce (invRev w) = invRev (reduce w) := by + apply reduce.min + rw [← red_invRev_iff, invRev_invRev] + apply Red.reduce_left + have : Red (invRev (invRev w)) (invRev (reduce (invRev w))) := reduce.red.invRev + rwa [invRev_invRev] at this + +@[to_additive (attr := simp)] +theorem toWord_inv (x : FreeGroup α) : x⁻¹.toWord = invRev x.toWord := by + rcases x with ⟨L⟩ + rw [quot_mk_eq_mk, inv_mk, toWord_mk, toWord_mk, reduce_invRev] + +open List -- for <+ notation + +@[to_additive] +lemma toWord_mul_sublist (x y : FreeGroup α) : (x * y).toWord <+ x.toWord ++ y.toWord := by + refine Red.sublist ?_ + have : x * y = FreeGroup.mk (x.toWord ++ y.toWord) := by + rw [← FreeGroup.mul_mk, FreeGroup.mk_toWord, FreeGroup.mk_toWord] + rw [this] + exact FreeGroup.reduce.red + +/-- **Constructive Church-Rosser theorem** (compare `church_rosser`). -/ +@[to_additive "**Constructive Church-Rosser theorem** (compare `church_rosser`)."] +def reduce.churchRosser (H12 : Red L₁ L₂) (H13 : Red L₁ L₃) : { L₄ // Red L₂ L₄ ∧ Red L₃ L₄ } := + ⟨reduce L₁, reduce.rev H12, reduce.rev H13⟩ + +@[to_additive] +instance : DecidableEq (FreeGroup α) := + toWord_injective.decidableEq + +-- TODO @[to_additive] doesn't succeed, possibly due to a bug +instance Red.decidableRel : DecidableRel (@Red α) + | [], [] => isTrue Red.refl + | [], _hd2 :: _tl2 => isFalse fun H => List.noConfusion (Red.nil_iff.1 H) + | (x, b) :: tl, [] => + match Red.decidableRel tl [(x, not b)] with + | isTrue H => isTrue <| Red.trans (Red.cons_cons H) <| (@Red.Step.not _ [] [] _ _).to_red + | isFalse H => isFalse fun H2 => H <| Red.cons_nil_iff_singleton.1 H2 + | (x1, b1) :: tl1, (x2, b2) :: tl2 => + if h : (x1, b1) = (x2, b2) then + match Red.decidableRel tl1 tl2 with + | isTrue H => isTrue <| h ▸ Red.cons_cons H + | isFalse H => isFalse fun H2 => H <| (Red.cons_cons_iff _).1 <| h.symm ▸ H2 + else + match Red.decidableRel tl1 ((x1, ! b1) :: (x2, b2) :: tl2) with + | isTrue H => isTrue <| (Red.cons_cons H).tail Red.Step.cons_not + | isFalse H => isFalse fun H2 => H <| Red.inv_of_red_of_ne h H2 + +/-- A list containing every word that `w₁` reduces to. -/ +def Red.enum (L₁ : List (α × Bool)) : List (List (α × Bool)) := + List.filter (Red L₁) (List.sublists L₁) + +theorem Red.enum.sound (H : L₂ ∈ List.filter (Red L₁) (List.sublists L₁)) : Red L₁ L₂ := + of_decide_eq_true (@List.of_mem_filter _ _ L₂ _ H) + +theorem Red.enum.complete (H : Red L₁ L₂) : L₂ ∈ Red.enum L₁ := + List.mem_filter_of_mem (List.mem_sublists.2 <| Red.sublist H) (decide_eq_true H) + +instance (L₁ : List (α × Bool)) : Fintype { L₂ // Red L₁ L₂ } := + Fintype.subtype (List.toFinset <| Red.enum L₁) fun _L₂ => + ⟨fun H => Red.enum.sound <| List.mem_toFinset.1 H, fun H => + List.mem_toFinset.2 <| Red.enum.complete H⟩ + +end Reduce + +@[to_additive (attr := simp)] +theorem one_ne_of (a : α) : 1 ≠ of a := + letI := Classical.decEq α; ne_of_apply_ne toWord <| by simp + +@[to_additive (attr := simp)] +theorem of_ne_one (a : α) : of a ≠ 1 := one_ne_of _ |>.symm + +@[to_additive] +instance [Nonempty α] : Nontrivial (FreeGroup α) where + exists_pair_ne := let ⟨x⟩ := ‹Nonempty α›; ⟨1, of x, one_ne_of x⟩ + +section Metric + +variable [DecidableEq α] + +/-- The length of reduced words provides a norm on a free group. -/ +@[to_additive "The length of reduced words provides a norm on an additive free group."] +def norm (x : FreeGroup α) : ℕ := + x.toWord.length + +@[to_additive (attr := simp)] +theorem norm_inv_eq {x : FreeGroup α} : norm x⁻¹ = norm x := by + simp only [norm, toWord_inv, invRev_length] + +@[to_additive (attr := simp)] +theorem norm_eq_zero {x : FreeGroup α} : norm x = 0 ↔ x = 1 := by + simp only [norm, List.length_eq_zero, toWord_eq_nil_iff] + +@[to_additive (attr := simp)] +theorem norm_one : norm (1 : FreeGroup α) = 0 := + rfl + +@[to_additive (attr := simp)] +theorem norm_of (a : α) : norm (of a) = 1 := + rfl + +@[to_additive] +theorem norm_mk_le : norm (mk L₁) ≤ L₁.length := + reduce.red.length_le + +@[to_additive] +theorem norm_mul_le (x y : FreeGroup α) : norm (x * y) ≤ norm x + norm y := + calc + norm (x * y) = norm (mk (x.toWord ++ y.toWord)) := by rw [← mul_mk, mk_toWord, mk_toWord] + _ ≤ (x.toWord ++ y.toWord).length := norm_mk_le + _ = norm x + norm y := List.length_append _ _ + +@[to_additive (attr := simp)] +theorem norm_of_pow (a : α) (n : ℕ) : norm (of a ^ n) = n := by + rw [norm, toWord_of_pow, List.length_replicate] + +@[to_additive] +theorem norm_surjective [Nonempty α] : Function.Surjective (norm (α := α)) := by + let ⟨a⟩ := ‹Nonempty α› + exact Function.RightInverse.surjective <| norm_of_pow a + +end Metric + +end FreeGroup diff --git a/Mathlib/SetTheory/Cardinal/Free.lean b/Mathlib/SetTheory/Cardinal/Free.lean index d802c1a3f9335..074fafe8ae1f3 100644 --- a/Mathlib/SetTheory/Cardinal/Free.lean +++ b/Mathlib/SetTheory/Cardinal/Free.lean @@ -5,6 +5,7 @@ Authors: Eric Wieser, Daniel Weber -/ import Mathlib.Data.Finsupp.Fintype import Mathlib.GroupTheory.FreeAbelianGroupFinsupp +import Mathlib.GroupTheory.FreeGroup.Reduce import Mathlib.RingTheory.FreeCommRing import Mathlib.SetTheory.Cardinal.Arithmetic import Mathlib.SetTheory.Cardinal.Finsupp From ebd1b1af5c41902abbe4beb77dc8c7db0c32f209 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Wed, 27 Nov 2024 09:56:45 +0000 Subject: [PATCH 345/829] =?UTF-8?q?feat(Algebra):=20ULift=20=E2=84=A4=20is?= =?UTF-8?q?=20initial=20in=20the=20category=20of=20commutative=20rings=20(?= =?UTF-8?q?#19507)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/Algebra/Category/Ring/Constructions.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/Mathlib/Algebra/Category/Ring/Constructions.lean b/Mathlib/Algebra/Category/Ring/Constructions.lean index ca21d51022790..c124625d6a67f 100644 --- a/Mathlib/Algebra/Category/Ring/Constructions.lean +++ b/Mathlib/Algebra/Category/Ring/Constructions.lean @@ -144,6 +144,14 @@ theorem subsingleton_of_isTerminal {X : CommRingCat} (hX : IsTerminal X) : Subsi def zIsInitial : IsInitial (CommRingCat.of ℤ) := IsInitial.ofUnique (h := fun R => ⟨⟨Int.castRingHom R⟩, fun a => a.ext_int _⟩) +/-- `ULift.{u} ℤ` is initial in `CommRingCat`. -/ +def isInitial : IsInitial (CommRingCat.of (ULift.{u} ℤ)) := + IsInitial.ofUnique (h := fun R ↦ ⟨⟨(Int.castRingHom R).comp ULift.ringEquiv.toRingHom⟩, + fun _ ↦ by + rw [← RingHom.cancel_right (f := (ULift.ringEquiv.{0, u} (α := ℤ)).symm.toRingHom) + (hf := ULift.ringEquiv.symm.surjective)] + apply RingHom.ext_int⟩) + end Terminal section Product From 3942daa810bd06acd4daefa3fc3771378c5b2acb Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Wed, 27 Nov 2024 09:56:46 +0000 Subject: [PATCH 346/829] chore(CategoryTheory/Sites): split DenseSubsite (#19519) The file `CategoryTheory.Sites.DenseSubsite` is split into `DenseSubsite.Basic` and `DenseSubsite.SheafEquiv`, the latter containing the equivalence of categories of sheaves in case suitable limits exists. The file `Sites.InducedTopology` is also moved to `Sites.DenseSubsite`. (This split prepares for #19444 where the equivalence of categories of sheaves is obtained under a somewhat weaker condition.) --- Mathlib.lean | 5 +- .../Sites/Coherent/SheafComparison.lean | 2 +- .../CategoryTheory/Sites/ConstantSheaf.lean | 2 +- .../Basic.lean} | 123 --------------- .../{ => DenseSubsite}/InducedTopology.lean | 2 +- .../Sites/DenseSubsite/SheafEquiv.lean | 142 ++++++++++++++++++ Mathlib/CategoryTheory/Sites/Equivalence.lean | 2 +- .../Sites/PreservesLocallyBijective.lean | 2 +- .../Sheaves/SheafCondition/Sites.lean | 2 +- 9 files changed, 151 insertions(+), 131 deletions(-) rename Mathlib/CategoryTheory/Sites/{DenseSubsite.lean => DenseSubsite/Basic.lean} (78%) rename Mathlib/CategoryTheory/Sites/{ => DenseSubsite}/InducedTopology.lean (98%) create mode 100644 Mathlib/CategoryTheory/Sites/DenseSubsite/SheafEquiv.lean diff --git a/Mathlib.lean b/Mathlib.lean index ff6ae6de6f371..521d0acd48e02 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2073,13 +2073,14 @@ import Mathlib.CategoryTheory.Sites.CoverLifting import Mathlib.CategoryTheory.Sites.CoverPreserving import Mathlib.CategoryTheory.Sites.Coverage import Mathlib.CategoryTheory.Sites.CoversTop -import Mathlib.CategoryTheory.Sites.DenseSubsite +import Mathlib.CategoryTheory.Sites.DenseSubsite.Basic +import Mathlib.CategoryTheory.Sites.DenseSubsite.InducedTopology +import Mathlib.CategoryTheory.Sites.DenseSubsite.SheafEquiv import Mathlib.CategoryTheory.Sites.EffectiveEpimorphic import Mathlib.CategoryTheory.Sites.EpiMono import Mathlib.CategoryTheory.Sites.EqualizerSheafCondition import Mathlib.CategoryTheory.Sites.Equivalence import Mathlib.CategoryTheory.Sites.Grothendieck -import Mathlib.CategoryTheory.Sites.InducedTopology import Mathlib.CategoryTheory.Sites.IsSheafFor import Mathlib.CategoryTheory.Sites.IsSheafOneHypercover import Mathlib.CategoryTheory.Sites.LeftExact diff --git a/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean b/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean index e78d9e4dc5cb6..5137ac167f17e 100644 --- a/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean +++ b/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean @@ -7,7 +7,7 @@ import Mathlib.CategoryTheory.Sites.Coherent.Comparison import Mathlib.CategoryTheory.Sites.Coherent.ExtensiveSheaves import Mathlib.CategoryTheory.Sites.Coherent.ReflectsPrecoherent import Mathlib.CategoryTheory.Sites.Coherent.ReflectsPreregular -import Mathlib.CategoryTheory.Sites.InducedTopology +import Mathlib.CategoryTheory.Sites.DenseSubsite.InducedTopology import Mathlib.CategoryTheory.Sites.Whiskering /-! diff --git a/Mathlib/CategoryTheory/Sites/ConstantSheaf.lean b/Mathlib/CategoryTheory/Sites/ConstantSheaf.lean index a300122560b4c..9bc3104f4e593 100644 --- a/Mathlib/CategoryTheory/Sites/ConstantSheaf.lean +++ b/Mathlib/CategoryTheory/Sites/ConstantSheaf.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Dagur Asgeirsson -/ import Mathlib.CategoryTheory.Sites.Sheafification -import Mathlib.CategoryTheory.Sites.DenseSubsite +import Mathlib.CategoryTheory.Sites.DenseSubsite.SheafEquiv /-! # The constant sheaf diff --git a/Mathlib/CategoryTheory/Sites/DenseSubsite.lean b/Mathlib/CategoryTheory/Sites/DenseSubsite/Basic.lean similarity index 78% rename from Mathlib/CategoryTheory/Sites/DenseSubsite.lean rename to Mathlib/CategoryTheory/Sites/DenseSubsite/Basic.lean index 5e24ab6fe6ff6..a7b6c780ed15f 100644 --- a/Mathlib/CategoryTheory/Sites/DenseSubsite.lean +++ b/Mathlib/CategoryTheory/Sites/DenseSubsite/Basic.lean @@ -29,9 +29,6 @@ that factors through images of the functor for each object in `D`. - `CategoryTheory.Functor.IsDenseSubsite`: The functor `G : C ⥤ D` exhibits `(C, J)` as a dense subsite of `(D, K)` if `G` is cover-dense, locally fully-faithful, and `S` is a cover of `C` iff the image of `S` in `D` is a cover. -- `CategoryTheory.Functor.IsDenseSubsite.sheafEquiv`: - If `G : C ⥤ D` exhibits `(C, J)` as a dense subsite of `(D, K)`, - it induces an equivalence of category of sheaves valued in a category with suitable limits. ## References @@ -41,7 +38,6 @@ that factors through images of the functor for each object in `D`. -/ - universe w v u namespace CategoryTheory @@ -556,122 +552,3 @@ end IsDenseSubsite end Functor end CategoryTheory - -namespace CategoryTheory.Functor.IsDenseSubsite - -open CategoryTheory Opposite - -universe w' -variable {C D : Type*} [Category C] [Category D] -variable (G : C ⥤ D) -variable (J : GrothendieckTopology C) (K : GrothendieckTopology D) -variable {A : Type w} [Category.{w'} A] [∀ X, Limits.HasLimitsOfShape (StructuredArrow X G.op) A] -variable [G.IsDenseSubsite J K] - -include K in -lemma isIso_ranCounit_app_of_isDenseSubsite (Y : Sheaf J A) (U X) : - IsIso ((yoneda.map ((G.op.ranCounit.app Y.val).app (op U))).app (op X)) := by - rw [isIso_iff_bijective] - constructor - · intro f₁ f₂ e - apply (isPointwiseRightKanExtensionRanCounit G.op Y.1 (.op (G.obj U))).hom_ext - rintro ⟨⟨⟨⟩⟩, ⟨W⟩, g⟩ - obtain ⟨g, rfl⟩ : ∃ g' : G.obj W ⟶ G.obj U, g = g'.op := ⟨g.unop, rfl⟩ - apply (Y.2 X _ (IsDenseSubsite.imageSieve_mem J K G g)).isSeparatedFor.ext - dsimp - rintro V iVW ⟨iVU, e'⟩ - have := congr($e ≫ Y.1.map iVU.op) - simp only [comp_obj, yoneda_map_app, Category.assoc, coyoneda_obj_obj, comp_map, - coyoneda_obj_map, ← NatTrans.naturality, op_obj, op_map, Quiver.Hom.unop_op, ← map_comp_assoc, - ← op_comp, ← e'] at this ⊢ - erw [← NatTrans.naturality] at this - exact this - · intro f - have (X Y Z) (f : X ⟶ Y) (g : G.obj Y ⟶ G.obj Z) (hf : G.imageSieve g f) : Exists _ := hf - choose l hl using this - let c : Limits.Cone (StructuredArrow.proj (op (G.obj U)) G.op ⋙ Y.val) := by - refine ⟨X, ⟨fun g ↦ ?_, ?_⟩⟩ - · refine Y.2.amalgamate ⟨_, IsDenseSubsite.imageSieve_mem J K G g.hom.unop⟩ - (fun I ↦ f ≫ Y.1.map (l _ _ _ _ _ I.hf).op) fun I₁ I₂ r ↦ ?_ - apply (Y.2 X _ (IsDenseSubsite.equalizer_mem J K G (r.g₁ ≫ l _ _ _ _ _ I₁.hf) - (r.g₂ ≫ l _ _ _ _ _ I₂.hf) ?_)).isSeparatedFor.ext fun V iUV (hiUV : _ = _) ↦ ?_ - · simp only [const_obj_obj, op_obj, map_comp, hl] - simp only [← map_comp_assoc, r.w] - · simp [← map_comp, ← op_comp, hiUV] - · dsimp - rintro ⟨⟨⟨⟩⟩, ⟨W₁⟩, g₁⟩ ⟨⟨⟨⟩⟩, ⟨W₂⟩, g₂⟩ ⟨⟨⟨⟨⟩⟩⟩, i, hi⟩ - dsimp at g₁ g₂ i hi - -- See issue https://github.com/leanprover-community/mathlib4/pull/15781 for tracking performance regressions of `rintro` as here - have h : g₂ = g₁ ≫ (G.map i.unop).op := by simpa only [Category.id_comp] using hi - rcases h with ⟨rfl⟩ - have h : ∃ g' : G.obj W₁ ⟶ G.obj U, g₁ = g'.op := ⟨g₁.unop, rfl⟩ - rcases h with ⟨g, rfl⟩ - have h : ∃ i' : W₂ ⟶ W₁, i = i'.op := ⟨i.unop, rfl⟩ - rcases h with ⟨i, rfl⟩ - simp only [const_obj_obj, id_obj, comp_obj, StructuredArrow.proj_obj, const_obj_map, op_obj, - unop_comp, Quiver.Hom.unop_op, Category.id_comp, comp_map, StructuredArrow.proj_map] - apply Y.2.hom_ext ⟨_, IsDenseSubsite.imageSieve_mem J K G (G.map i ≫ g)⟩ - intro I - simp only [Presheaf.IsSheaf.amalgamate_map, Category.assoc, ← Functor.map_comp, ← op_comp] - let I' : GrothendieckTopology.Cover.Arrow ⟨_, IsDenseSubsite.imageSieve_mem J K G g⟩ := - ⟨_, I.f ≫ i, ⟨l _ _ _ _ _ I.hf, by simp [hl]⟩⟩ - refine Eq.trans ?_ (Y.2.amalgamate_map _ _ _ I').symm - apply (Y.2 X _ (IsDenseSubsite.equalizer_mem J K G (l _ _ _ _ _ I.hf) - (l _ _ _ _ _ I'.hf) (by simp [hl]))).isSeparatedFor.ext fun V iUV (hiUV : _ = _) ↦ ?_ - simp [← Functor.map_comp, ← op_comp, hiUV] - refine ⟨(isPointwiseRightKanExtensionRanCounit G.op Y.1 (.op (G.obj U))).lift c, ?_⟩ - · have := (isPointwiseRightKanExtensionRanCounit G.op Y.1 (.op (G.obj U))).fac c (.mk (𝟙 _)) - simp only [id_obj, comp_obj, StructuredArrow.proj_obj, StructuredArrow.mk_right, - RightExtension.coneAt_pt, RightExtension.mk_left, RightExtension.coneAt_π_app, - const_obj_obj, op_obj, StructuredArrow.mk_hom_eq_self, map_id, whiskeringLeft_obj_obj, - RightExtension.mk_hom, Category.id_comp, StructuredArrow.mk_left, unop_id] at this - simp only [id_obj, yoneda_map_app, this] - apply Y.2.hom_ext ⟨_, IsDenseSubsite.imageSieve_mem J K G (𝟙 (G.obj U))⟩ _ _ fun I ↦ ?_ - apply (Y.2 X _ (IsDenseSubsite.equalizer_mem J K G (l _ _ _ _ _ I.hf) - I.f (by simp [hl]))).isSeparatedFor.ext fun V iUV (hiUV : _ = _) ↦ ?_ - simp [← Functor.map_comp, ← op_comp, hiUV] - -instance (Y : Sheaf J A) : IsIso ((G.sheafAdjunctionCocontinuous A J K).counit.app Y) := by - apply (config := { allowSynthFailures := true }) - ReflectsIsomorphisms.reflects (sheafToPresheaf J A) - rw [NatTrans.isIso_iff_isIso_app] - intro ⟨U⟩ - apply (config := { allowSynthFailures := true }) ReflectsIsomorphisms.reflects yoneda - rw [NatTrans.isIso_iff_isIso_app] - intro ⟨X⟩ - simp only [comp_obj, sheafToPresheaf_obj, sheafPushforwardContinuous_obj_val_obj, yoneda_obj_obj, - id_obj, sheafToPresheaf_map, sheafAdjunctionCocontinuous_counit_app_val, ranAdjunction_counit] - exact isIso_ranCounit_app_of_isDenseSubsite G J K Y U X - -variable (A) - -/-- -If `G : C ⥤ D` exhibits `(C, J)` as a dense subsite of `(D, K)`, -it induces an equivalence of category of sheaves valued in a category with suitable limits. --/ -@[simps! functor inverse] -noncomputable def sheafEquiv : Sheaf J A ≌ Sheaf K A := - (G.sheafAdjunctionCocontinuous A J K).toEquivalence.symm - -instance : (G.sheafPushforwardContinuous A J K).IsEquivalence := - inferInstanceAs (IsDenseSubsite.sheafEquiv G _ _ _).inverse.IsEquivalence - -variable [HasWeakSheafify J A] [HasWeakSheafify K A] - -/-- The natural isomorphism exhibiting the compatibility of -`IsDenseSubsite.sheafEquiv` with sheafification. -/ -noncomputable -abbrev sheafEquivSheafificationCompatibility : - (whiskeringLeft _ _ A).obj G.op ⋙ presheafToSheaf _ _ ≅ - presheafToSheaf _ _ ⋙ (sheafEquiv G J K A).inverse := by - apply Functor.pushforwardContinuousSheafificationCompatibility - -end IsDenseSubsite - -@[deprecated (since := "2024-07-23")] -alias IsCoverDense.sheafEquivOfCoverPreservingCoverLifting := IsDenseSubsite.sheafEquiv -@[deprecated (since := "2024-07-23")] -alias IsCoverDense.sheafEquivOfCoverPreservingCoverLiftingSheafificationCompatibility := - IsDenseSubsite.sheafEquivSheafificationCompatibility - -end CategoryTheory.Functor diff --git a/Mathlib/CategoryTheory/Sites/InducedTopology.lean b/Mathlib/CategoryTheory/Sites/DenseSubsite/InducedTopology.lean similarity index 98% rename from Mathlib/CategoryTheory/Sites/InducedTopology.lean rename to Mathlib/CategoryTheory/Sites/DenseSubsite/InducedTopology.lean index c5be7f7e1ffda..4710ae5d4bf61 100644 --- a/Mathlib/CategoryTheory/Sites/InducedTopology.lean +++ b/Mathlib/CategoryTheory/Sites/DenseSubsite/InducedTopology.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ -import Mathlib.CategoryTheory.Sites.DenseSubsite +import Mathlib.CategoryTheory.Sites.DenseSubsite.SheafEquiv /-! # Induced Topology diff --git a/Mathlib/CategoryTheory/Sites/DenseSubsite/SheafEquiv.lean b/Mathlib/CategoryTheory/Sites/DenseSubsite/SheafEquiv.lean new file mode 100644 index 0000000000000..7cc5b9aefa271 --- /dev/null +++ b/Mathlib/CategoryTheory/Sites/DenseSubsite/SheafEquiv.lean @@ -0,0 +1,142 @@ +/- +Copyright (c) 2021 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ + +import Mathlib.CategoryTheory.Sites.DenseSubsite.Basic + +/-! +# The equivalence of categories of sheaves of a dense subsite + +- `CategoryTheory.Functor.IsDenseSubsite.sheafEquiv`: + If `G : C ⥤ D` exhibits `(C, J)` as a dense subsite of `(D, K)`, + it induces an equivalence of category of sheaves valued in a category with suitable limits. + +## References + +* [Elephant]: *Sketches of an Elephant*, ℱ. T. Johnstone: C2.2. +* https://ncatlab.org/nlab/show/dense+sub-site +* https://ncatlab.org/nlab/show/comparison+lemma + +-/ + +universe w v u w' + +namespace CategoryTheory.Functor.IsDenseSubsite + +open CategoryTheory Opposite + +variable {C D : Type*} [Category C] [Category D] +variable (G : C ⥤ D) +variable (J : GrothendieckTopology C) (K : GrothendieckTopology D) +variable {A : Type w} [Category.{w'} A] [∀ X, Limits.HasLimitsOfShape (StructuredArrow X G.op) A] +variable [G.IsDenseSubsite J K] + +include K in +lemma isIso_ranCounit_app_of_isDenseSubsite (Y : Sheaf J A) (U X) : + IsIso ((yoneda.map ((G.op.ranCounit.app Y.val).app (op U))).app (op X)) := by + rw [isIso_iff_bijective] + constructor + · intro f₁ f₂ e + apply (isPointwiseRightKanExtensionRanCounit G.op Y.1 (.op (G.obj U))).hom_ext + rintro ⟨⟨⟨⟩⟩, ⟨W⟩, g⟩ + obtain ⟨g, rfl⟩ : ∃ g' : G.obj W ⟶ G.obj U, g = g'.op := ⟨g.unop, rfl⟩ + apply (Y.2 X _ (IsDenseSubsite.imageSieve_mem J K G g)).isSeparatedFor.ext + dsimp + rintro V iVW ⟨iVU, e'⟩ + have := congr($e ≫ Y.1.map iVU.op) + simp only [comp_obj, yoneda_map_app, Category.assoc, coyoneda_obj_obj, comp_map, + coyoneda_obj_map, ← NatTrans.naturality, op_obj, op_map, Quiver.Hom.unop_op, ← map_comp_assoc, + ← op_comp, ← e'] at this ⊢ + erw [← NatTrans.naturality] at this + exact this + · intro f + have (X Y Z) (f : X ⟶ Y) (g : G.obj Y ⟶ G.obj Z) (hf : G.imageSieve g f) : Exists _ := hf + choose l hl using this + let c : Limits.Cone (StructuredArrow.proj (op (G.obj U)) G.op ⋙ Y.val) := by + refine ⟨X, ⟨fun g ↦ ?_, ?_⟩⟩ + · refine Y.2.amalgamate ⟨_, IsDenseSubsite.imageSieve_mem J K G g.hom.unop⟩ + (fun I ↦ f ≫ Y.1.map (l _ _ _ _ _ I.hf).op) fun I₁ I₂ r ↦ ?_ + apply (Y.2 X _ (IsDenseSubsite.equalizer_mem J K G (r.g₁ ≫ l _ _ _ _ _ I₁.hf) + (r.g₂ ≫ l _ _ _ _ _ I₂.hf) ?_)).isSeparatedFor.ext fun V iUV (hiUV : _ = _) ↦ ?_ + · simp only [const_obj_obj, op_obj, map_comp, hl] + simp only [← map_comp_assoc, r.w] + · simp [← map_comp, ← op_comp, hiUV] + · dsimp + rintro ⟨⟨⟨⟩⟩, ⟨W₁⟩, g₁⟩ ⟨⟨⟨⟩⟩, ⟨W₂⟩, g₂⟩ ⟨⟨⟨⟨⟩⟩⟩, i, hi⟩ + dsimp at g₁ g₂ i hi + -- See issue https://github.com/leanprover-community/mathlib4/pull/15781 for tracking performance regressions of `rintro` as here + have h : g₂ = g₁ ≫ (G.map i.unop).op := by simpa only [Category.id_comp] using hi + rcases h with ⟨rfl⟩ + have h : ∃ g' : G.obj W₁ ⟶ G.obj U, g₁ = g'.op := ⟨g₁.unop, rfl⟩ + rcases h with ⟨g, rfl⟩ + have h : ∃ i' : W₂ ⟶ W₁, i = i'.op := ⟨i.unop, rfl⟩ + rcases h with ⟨i, rfl⟩ + simp only [const_obj_obj, id_obj, comp_obj, StructuredArrow.proj_obj, const_obj_map, op_obj, + unop_comp, Quiver.Hom.unop_op, Category.id_comp, comp_map, StructuredArrow.proj_map] + apply Y.2.hom_ext ⟨_, IsDenseSubsite.imageSieve_mem J K G (G.map i ≫ g)⟩ + intro I + simp only [Presheaf.IsSheaf.amalgamate_map, Category.assoc, ← Functor.map_comp, ← op_comp] + let I' : GrothendieckTopology.Cover.Arrow ⟨_, IsDenseSubsite.imageSieve_mem J K G g⟩ := + ⟨_, I.f ≫ i, ⟨l _ _ _ _ _ I.hf, by simp [hl]⟩⟩ + refine Eq.trans ?_ (Y.2.amalgamate_map _ _ _ I').symm + apply (Y.2 X _ (IsDenseSubsite.equalizer_mem J K G (l _ _ _ _ _ I.hf) + (l _ _ _ _ _ I'.hf) (by simp [hl]))).isSeparatedFor.ext fun V iUV (hiUV : _ = _) ↦ ?_ + simp [← Functor.map_comp, ← op_comp, hiUV] + refine ⟨(isPointwiseRightKanExtensionRanCounit G.op Y.1 (.op (G.obj U))).lift c, ?_⟩ + · have := (isPointwiseRightKanExtensionRanCounit G.op Y.1 (.op (G.obj U))).fac c (.mk (𝟙 _)) + simp only [id_obj, comp_obj, StructuredArrow.proj_obj, StructuredArrow.mk_right, + RightExtension.coneAt_pt, RightExtension.mk_left, RightExtension.coneAt_π_app, + const_obj_obj, op_obj, StructuredArrow.mk_hom_eq_self, map_id, whiskeringLeft_obj_obj, + RightExtension.mk_hom, Category.id_comp, StructuredArrow.mk_left, unop_id] at this + simp only [id_obj, yoneda_map_app, this] + apply Y.2.hom_ext ⟨_, IsDenseSubsite.imageSieve_mem J K G (𝟙 (G.obj U))⟩ _ _ fun I ↦ ?_ + apply (Y.2 X _ (IsDenseSubsite.equalizer_mem J K G (l _ _ _ _ _ I.hf) + I.f (by simp [hl]))).isSeparatedFor.ext fun V iUV (hiUV : _ = _) ↦ ?_ + simp [← Functor.map_comp, ← op_comp, hiUV] + +instance (Y : Sheaf J A) : IsIso ((G.sheafAdjunctionCocontinuous A J K).counit.app Y) := by + apply (config := { allowSynthFailures := true }) + ReflectsIsomorphisms.reflects (sheafToPresheaf J A) + rw [NatTrans.isIso_iff_isIso_app] + intro ⟨U⟩ + apply (config := { allowSynthFailures := true }) ReflectsIsomorphisms.reflects yoneda + rw [NatTrans.isIso_iff_isIso_app] + intro ⟨X⟩ + simp only [comp_obj, sheafToPresheaf_obj, sheafPushforwardContinuous_obj_val_obj, yoneda_obj_obj, + id_obj, sheafToPresheaf_map, sheafAdjunctionCocontinuous_counit_app_val, ranAdjunction_counit] + exact isIso_ranCounit_app_of_isDenseSubsite G J K Y U X + +variable (A) + +/-- +If `G : C ⥤ D` exhibits `(C, J)` as a dense subsite of `(D, K)`, +it induces an equivalence of category of sheaves valued in a category with suitable limits. +-/ +@[simps! functor inverse] +noncomputable def sheafEquiv : Sheaf J A ≌ Sheaf K A := + (G.sheafAdjunctionCocontinuous A J K).toEquivalence.symm + +instance : (G.sheafPushforwardContinuous A J K).IsEquivalence := + inferInstanceAs (IsDenseSubsite.sheafEquiv G _ _ _).inverse.IsEquivalence + +variable [HasWeakSheafify J A] [HasWeakSheafify K A] + +/-- The natural isomorphism exhibiting the compatibility of +`IsDenseSubsite.sheafEquiv` with sheafification. -/ +noncomputable +abbrev sheafEquivSheafificationCompatibility : + (whiskeringLeft _ _ A).obj G.op ⋙ presheafToSheaf _ _ ≅ + presheafToSheaf _ _ ⋙ (sheafEquiv G J K A).inverse := by + apply Functor.pushforwardContinuousSheafificationCompatibility + +end IsDenseSubsite + +@[deprecated (since := "2024-07-23")] +alias IsCoverDense.sheafEquivOfCoverPreservingCoverLifting := IsDenseSubsite.sheafEquiv +@[deprecated (since := "2024-07-23")] +alias IsCoverDense.sheafEquivOfCoverPreservingCoverLiftingSheafificationCompatibility := + IsDenseSubsite.sheafEquivSheafificationCompatibility + +end CategoryTheory.Functor diff --git a/Mathlib/CategoryTheory/Sites/Equivalence.lean b/Mathlib/CategoryTheory/Sites/Equivalence.lean index 6953348d2088a..9ae28dadabe3f 100644 --- a/Mathlib/CategoryTheory/Sites/Equivalence.lean +++ b/Mathlib/CategoryTheory/Sites/Equivalence.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Dagur Asgeirsson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Dagur Asgeirsson -/ -import Mathlib.CategoryTheory.Sites.InducedTopology +import Mathlib.CategoryTheory.Sites.DenseSubsite.InducedTopology import Mathlib.CategoryTheory.Sites.LocallyBijective import Mathlib.CategoryTheory.Sites.PreservesLocallyBijective import Mathlib.CategoryTheory.Sites.Whiskering diff --git a/Mathlib/CategoryTheory/Sites/PreservesLocallyBijective.lean b/Mathlib/CategoryTheory/Sites/PreservesLocallyBijective.lean index bbc801acd097d..ab499ecc7dfde 100644 --- a/Mathlib/CategoryTheory/Sites/PreservesLocallyBijective.lean +++ b/Mathlib/CategoryTheory/Sites/PreservesLocallyBijective.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Dagur Asgeirsson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Dagur Asgeirsson -/ -import Mathlib.CategoryTheory.Sites.DenseSubsite +import Mathlib.CategoryTheory.Sites.DenseSubsite.Basic import Mathlib.CategoryTheory.Sites.LocallySurjective /-! diff --git a/Mathlib/Topology/Sheaves/SheafCondition/Sites.lean b/Mathlib/Topology/Sheaves/SheafCondition/Sites.lean index d4c904a20ecf8..185316bdf4310 100644 --- a/Mathlib/Topology/Sheaves/SheafCondition/Sites.lean +++ b/Mathlib/Topology/Sheaves/SheafCondition/Sites.lean @@ -5,7 +5,7 @@ Authors: Justus Springer -/ import Mathlib.CategoryTheory.Sites.Spaces import Mathlib.Topology.Sheaves.Sheaf -import Mathlib.CategoryTheory.Sites.DenseSubsite +import Mathlib.CategoryTheory.Sites.DenseSubsite.Basic /-! From cfef5c4f0a57295f55f26a41cf24ae4d5ca2b613 Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Wed, 27 Nov 2024 10:10:36 +0000 Subject: [PATCH 347/829] doc: clarify doc of HasProd (#19530) --- Mathlib/Topology/Algebra/InfiniteSum/Defs.lean | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean b/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean index cc83507487809..44ab58cf88b01 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean @@ -59,7 +59,8 @@ section HasProd variable [CommMonoid α] [TopologicalSpace α] -/-- Infinite product on a topological monoid +/-- `HasProd f a` means that the (potentially infinite) product of the `f b` for `b : β` converges +to `a`. The `atTop` filter on `Finset β` is the limit of all finite sets towards the entire type. So we take the product over bigger and bigger sets. This product operation is invariant under reordering. @@ -70,7 +71,8 @@ this assumption later, for the lemmas where it is relevant. These are defined in an identical way to infinite sums (`HasSum`). For example, we say that the function `ℕ → ℝ` sending `n` to `1 / 2` has a product of `0`, rather than saying that it does not converge as some authors would. -/ -@[to_additive "Infinite sum on a topological monoid +@[to_additive "`HasSum f a` means that the (potentially infinite) sum of the `f b` for `b : β` +converges to `a`. The `atTop` filter on `Finset β` is the limit of all finite sets towards the entire type. So we sum up bigger and bigger sets. This sum operation is invariant under reordering. In particular, From c59545b6f988c2439d89047e58cef992b0dfa62a Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Wed, 27 Nov 2024 10:31:12 +0000 Subject: [PATCH 348/829] feat: ind-objects are closed under products (#18988) Co-authored-by: Markus Himmel --- Mathlib.lean | 2 + .../FilteredColimitCommutesProduct.lean | 185 ++++++++++++++++++ .../Limits/Indization/Products.lean | 43 ++++ 3 files changed, 230 insertions(+) create mode 100644 Mathlib/CategoryTheory/Limits/FilteredColimitCommutesProduct.lean create mode 100644 Mathlib/CategoryTheory/Limits/Indization/Products.lean diff --git a/Mathlib.lean b/Mathlib.lean index 521d0acd48e02..f3e568ce173a2 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1797,6 +1797,7 @@ import Mathlib.CategoryTheory.Limits.EssentiallySmall import Mathlib.CategoryTheory.Limits.ExactFunctor import Mathlib.CategoryTheory.Limits.Filtered import Mathlib.CategoryTheory.Limits.FilteredColimitCommutesFiniteLimit +import Mathlib.CategoryTheory.Limits.FilteredColimitCommutesProduct import Mathlib.CategoryTheory.Limits.Final import Mathlib.CategoryTheory.Limits.Final.ParallelPair import Mathlib.CategoryTheory.Limits.FinallySmall @@ -1816,6 +1817,7 @@ import Mathlib.CategoryTheory.Limits.Indization.Equalizers import Mathlib.CategoryTheory.Limits.Indization.FilteredColimits import Mathlib.CategoryTheory.Limits.Indization.IndObject import Mathlib.CategoryTheory.Limits.Indization.LocallySmall +import Mathlib.CategoryTheory.Limits.Indization.Products import Mathlib.CategoryTheory.Limits.IsConnected import Mathlib.CategoryTheory.Limits.IsLimit import Mathlib.CategoryTheory.Limits.Lattice diff --git a/Mathlib/CategoryTheory/Limits/FilteredColimitCommutesProduct.lean b/Mathlib/CategoryTheory/Limits/FilteredColimitCommutesProduct.lean new file mode 100644 index 0000000000000..ac195d5acac16 --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/FilteredColimitCommutesProduct.lean @@ -0,0 +1,185 @@ +/- +Copyright (c) 2024 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel +-/ +import Mathlib.CategoryTheory.Limits.FunctorCategory.Filtered +import Mathlib.CategoryTheory.Limits.Shapes.Types +import Mathlib.CategoryTheory.Limits.TypesFiltered +import Mathlib.CategoryTheory.Limits.FunctorCategory.Shapes.Products + +/-! +# The IPC property + +Given a family of categories `I i` (`i : α`) and a family of functors `F i : I i ⥤ C`, we construct +the natural morphism `colim_k (∏ᶜ s ↦ (F s).obj (k s)) ⟶ ∏ᶜ s ↦ colim_k (F s).obj (k s)`. + +Similarly to the study of finite limits commuting with filtered colimits, we then study sufficient +conditions for this morphism to be an isomorphism. We say that `C` satisfies the `w`-IPC property if +the morphism is an isomorphism as long as `α` is `w`-small and `I i` is `w`-small and filtered for +all `i`. + +We show that +* the category `Type u` satisfies the `u`-IPC property and +* if `C` satisfies the `w`-IPC property, then `D ⥤ C` satisfies the `w`-IPC property. + +These results will be used to show that if a category `C` has products indexed by `α`, then so +does the category of Ind-objects of `C`. + +## References +* [M. Kashiwara, P. Schapira, *Categories and Sheaves*][Kashiwara2006], 3.1.10, 3.1.11, 3.1.12. +-/ + +universe w v v₁ v₂ u u₁ u₂ + +namespace CategoryTheory.Limits + +section + +variable {C : Type u} [Category.{v} C] {α : Type w} {I : α → Type u₁} [∀ i, Category.{v₁} (I i)] + [HasLimitsOfShape (Discrete α) C] + (F : ∀ i, I i ⥤ C) + +/-- Given a family of functors `I i ⥤ C` for `i : α`, we obtain a functor `(∀ i, I i) ⥤ C` which +maps `k : ∀ i, I i` to `∏ᶜ fun (s : α) => (F s).obj (k s)`. -/ +@[simps] +noncomputable def pointwiseProduct : (∀ i, I i) ⥤ C where + obj k := ∏ᶜ fun (s : α) => (F s).obj (k s) + map f := Pi.map (fun s => (F s).map (f s)) + +variable [∀ i, HasColimitsOfShape (I i) C] [HasColimitsOfShape (∀ i, I i) C] + +/-- The inclusions `(F s).obj (k s) ⟶ colimit (F s)` induce a cocone on `pointwiseProduct F` with +cone point `∏ᶜ (fun s : α) => colimit (F s)`. -/ +@[simps] +noncomputable def coconePointwiseProduct : Cocone (pointwiseProduct F) where + pt := ∏ᶜ fun (s : α) => colimit (F s) + ι := { app := fun k => Pi.map fun s => colimit.ι _ _ } + +/-- The natural morphism `colim_k (∏ᶜ s ↦ (F s).obj (k s)) ⟶ ∏ᶜ s ↦ colim_k (F s).obj (k s)`. +We will say that a category has the `IPC` property if this morphism is an isomorphism as long +as the indexing categories are filtered. -/ +noncomputable def colimitPointwiseProductToProductColimit : + colimit (pointwiseProduct F) ⟶ ∏ᶜ fun (s : α) => colimit (F s) := + colimit.desc (pointwiseProduct F) (coconePointwiseProduct F) + +@[reassoc (attr := simp)] +theorem ι_colimitPointwiseProductToProductColimit_π (k : ∀ i, I i) (s : α) : + colimit.ι (pointwiseProduct F) k ≫ colimitPointwiseProductToProductColimit F ≫ Pi.π _ s = + Pi.π _ s ≫ colimit.ι (F s) (k s) := by + simp [colimitPointwiseProductToProductColimit] + +end + +section functorCategory + +variable {C : Type u} [Category.{v} C] {D : Type u₁} [Category.{v₁} D] + {α : Type w} {I : α → Type u₂} [∀ i, Category (I i)] + [HasLimitsOfShape (Discrete α) C] + (F : ∀ i, I i ⥤ D ⥤ C) + +/-- Evaluating the pointwise product `k ↦ ∏ᶜ fun (s : α) => (F s).obj (k s)` at `d` is the same as +taking the pointwise product `k ↦ ∏ᶜ fun (s : α) => ((F s).obj (k s)).obj d`. -/ +@[simps!] +noncomputable def pointwiseProductCompEvaluation (d : D) : + pointwiseProduct F ⋙ (evaluation D C).obj d ≅ + pointwiseProduct (fun s => F s ⋙ (evaluation _ _).obj d) := + NatIso.ofComponents (fun k => piObjIso _ _) + (fun f => Pi.hom_ext _ _ (by simp [← NatTrans.comp_app])) + +variable [∀ i, HasColimitsOfShape (I i) C] [HasColimitsOfShape (∀ i, I i) C] + +theorem colimitPointwiseProductToProductColimit_app (d : D) : + (colimitPointwiseProductToProductColimit F).app d = + (colimitObjIsoColimitCompEvaluation _ _).hom ≫ + (HasColimit.isoOfNatIso (pointwiseProductCompEvaluation F d)).hom ≫ + colimitPointwiseProductToProductColimit _ ≫ + (Pi.mapIso fun _ => (colimitObjIsoColimitCompEvaluation _ _).symm).hom ≫ + (piObjIso _ _).inv := by + rw [← Iso.inv_comp_eq] + simp only [← Category.assoc] + rw [Iso.eq_comp_inv] + refine Pi.hom_ext _ _ (fun s => colimit.hom_ext (fun k => ?_)) + simp [← NatTrans.comp_app] + +end functorCategory + +section + +variable (C : Type u) [Category.{v} C] + +/-- A category `C` has the `w`-IPC property if the natural morphism +`colim_k (∏ᶜ s ↦ (F s).obj (k s)) ⟶ ∏ᶜ s ↦ colim_k (F s).obj (k s)` is an isomorphism for any +family of functors `F i : I i ⥤ C` with `I i` `w`-small and filtered for all `i`. -/ +class IsIPC [HasProducts.{w} C] [HasFilteredColimitsOfSize.{w} C] : Prop where + /-- `colimitPointwiseProductToProductColimit F` is always an isomorphism. -/ + isIso : ∀ (α : Type w) (I : α → Type w) [∀ i, SmallCategory (I i)] [∀ i, IsFiltered (I i)] + (F : ∀ i, I i ⥤ C), IsIso (colimitPointwiseProductToProductColimit F) + +attribute [instance] IsIPC.isIso + +end + +section types + +variable {α : Type u} {I : α → Type u} [∀ i, SmallCategory (I i)] [∀ i, IsFiltered (I i)] + +theorem Types.isIso_colimitPointwiseProductToProductColimit (F : ∀ i, I i ⥤ Type u) : + IsIso (colimitPointwiseProductToProductColimit F) := by + -- We follow the proof in [Kashiwara2006], Prop. 3.1.11(ii) + refine (isIso_iff_bijective _).2 ⟨fun y y' hy => ?_, fun x => ?_⟩ + · obtain ⟨ky, yk₀, hyk₀⟩ := Types.jointly_surjective' y + obtain ⟨ky', yk₀', hyk₀'⟩ := Types.jointly_surjective' y' + let k := IsFiltered.max ky ky' + let yk : (pointwiseProduct F).obj k := + (pointwiseProduct F).map (IsFiltered.leftToMax ky ky') yk₀ + let yk' : (pointwiseProduct F).obj k := + (pointwiseProduct F).map (IsFiltered.rightToMax ky ky') yk₀' + obtain rfl : y = colimit.ι (pointwiseProduct F) k yk := by + simp only [yk, Types.Colimit.w_apply', hyk₀] + obtain rfl : y' = colimit.ι (pointwiseProduct F) k yk' := by + simp only [yk', Types.Colimit.w_apply', hyk₀'] + dsimp only [pointwiseProduct_obj] at yk yk' + have hch : ∀ (s : α), ∃ (i' : I s) (hi' : k s ⟶ i'), + (F s).map hi' (Pi.π (fun s => (F s).obj (k s)) s yk) = + (F s).map hi' (Pi.π (fun s => (F s).obj (k s)) s yk') := by + intro s + have hy₁ := congrFun (ι_colimitPointwiseProductToProductColimit_π F k s) yk + have hy₂ := congrFun (ι_colimitPointwiseProductToProductColimit_π F k s) yk' + dsimp only [pointwiseProduct_obj, types_comp_apply] at hy₁ hy₂ + rw [← hy, hy₁, Types.FilteredColimit.colimit_eq_iff] at hy₂ + obtain ⟨i₀, f₀, g₀, h₀⟩ := hy₂ + refine ⟨IsFiltered.coeq f₀ g₀, f₀ ≫ IsFiltered.coeqHom f₀ g₀, ?_⟩ + conv_rhs => rw [IsFiltered.coeq_condition] + simp [h₀] + choose k' f hk' using hch + apply Types.colimit_sound' f f + exact Types.limit_ext' _ _ _ (fun ⟨s⟩ => by simpa using hk' _) + · have hch : ∀ (s : α), ∃ (i : I s) (xi : (F s).obj i), colimit.ι (F s) i xi = + Pi.π (fun s => colimit (F s)) s x := fun s => Types.jointly_surjective' _ + choose k p hk using hch + refine ⟨colimit.ι (pointwiseProduct F) k ((Types.productIso _).inv p), ?_⟩ + refine Types.limit_ext' _ _ _ (fun ⟨s⟩ => ?_) + have := congrFun (ι_colimitPointwiseProductToProductColimit_π F k s) + ((Types.productIso _).inv p) + exact this.trans (by simpa using hk _) + +instance : IsIPC.{u} (Type u) where + isIso _ _ := Types.isIso_colimitPointwiseProductToProductColimit + +end types + +section functorCategory + +variable {C : Type u} [Category.{v} C] + +instance [HasProducts.{w} C] [HasFilteredColimitsOfSize.{w, w} C] [IsIPC.{w} C] {D : Type u₁} + [Category.{v₁} D] : IsIPC.{w} (D ⥤ C) := by + refine ⟨fun β I _ _ F => ?_⟩ + suffices ∀ d, IsIso ((colimitPointwiseProductToProductColimit F).app d) from + NatIso.isIso_of_isIso_app _ + exact fun d => colimitPointwiseProductToProductColimit_app F d ▸ inferInstance + +end functorCategory + +end CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Limits/Indization/Products.lean b/Mathlib/CategoryTheory/Limits/Indization/Products.lean new file mode 100644 index 0000000000000..ba9c97fd879ee --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Indization/Products.lean @@ -0,0 +1,43 @@ +/- +Copyright (c) 2024 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel +-/ +import Mathlib.CategoryTheory.Limits.FilteredColimitCommutesProduct +import Mathlib.CategoryTheory.Limits.Indization.FilteredColimits + +/-! +# Ind-objects are closed under products + +We show that if `C` admits products indexed by `α`, then `IsIndObject` is closed under taking +products in `Cᵒᵖ ⥤ Type v` indexed by `α`. This will imply that the functor `Ind C ⥤ Cᵒᵖ ⥤ Type v` +creates products indexed by `α` and that the functor `C ⥤ Ind C` preserves them. + +## References +* [M. Kashiwara, P. Schapira, *Categories and Sheaves*][Kashiwara2006], Prop. 6.1.16(ii) +-/ + +universe v u + +namespace CategoryTheory.Limits + +variable {C : Type u} [Category.{v} C] {α : Type v} + +theorem isIndObject_pi (h : ∀ (g : α → C), IsIndObject (∏ᶜ yoneda.obj ∘ g)) + (f : α → Cᵒᵖ ⥤ Type v) (hf : ∀ a, IsIndObject (f a)) : IsIndObject (∏ᶜ f) := by + let F := fun a => (hf a).presentation.F ⋙ yoneda + suffices (∏ᶜ f ≅ colimit (pointwiseProduct F)) from + (isIndObject_colimit _ _ (fun i => h _)).map this.inv + refine Pi.mapIso (fun s => ?_) ≪≫ (asIso (colimitPointwiseProductToProductColimit F)).symm + exact IsColimit.coconePointUniqueUpToIso (hf s).presentation.isColimit (colimit.isColimit _) + +theorem isIndObject_limit_of_discrete (h : ∀ (g : α → C), IsIndObject (∏ᶜ yoneda.obj ∘ g)) + (F : Discrete α ⥤ Cᵒᵖ ⥤ Type v) (hF : ∀ a, IsIndObject (F.obj a)) : IsIndObject (limit F) := + IsIndObject.map (Pi.isoLimit _).hom (isIndObject_pi h _ (fun a => hF ⟨a⟩)) + +theorem isIndObject_limit_of_discrete_of_hasLimitsOfShape [HasLimitsOfShape (Discrete α) C] + (F : Discrete α ⥤ Cᵒᵖ ⥤ Type v) (hF : ∀ a, IsIndObject (F.obj a)) : IsIndObject (limit F) := + isIndObject_limit_of_discrete (fun g => (isIndObject_limit_comp_yoneda (Discrete.functor g)).map + (HasLimit.isoOfNatIso (Discrete.compNatIsoDiscrete g yoneda)).hom) F hF + +end CategoryTheory.Limits From 3bcc7678efce412065e09ed91f8040020ff052ae Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 27 Nov 2024 10:56:04 +0000 Subject: [PATCH 349/829] chore(zulip_emoji_merge_delegate): exclude #rss channel from search (#19533) --- scripts/zulip_emoji_merge_delegate.py | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/scripts/zulip_emoji_merge_delegate.py b/scripts/zulip_emoji_merge_delegate.py index 8fa4e33e74f4c..cbecbdf203c10 100644 --- a/scripts/zulip_emoji_merge_delegate.py +++ b/scripts/zulip_emoji_merge_delegate.py @@ -23,14 +23,14 @@ site=ZULIP_SITE ) -# Fetch the messages containing the PR number from the public channels. +# Fetch the messages containing the PR number from the public channels (exluding #rss). # There does not seem to be a way to search simultaneously public and private channels. public_response = client.get_messages({ "anchor": "newest", "num_before": 5000, "num_after": 0, "narrow": [ - {"operator": "channels", "operand": "public"}, + {"operator": "channel", "operand": "rss", "negated": true}, {"operator": "search", "operand": f'#{PR_NUMBER}'}, ], }) From 01b31e1d1d58001128368257a32bd3337a5b576e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 27 Nov 2024 11:08:08 +0000 Subject: [PATCH 350/829] chore(Submonoid/Subgroup): rename `subtype_range` to `range_subtype` (#19518) This matches the naming convention. --- Mathlib/Algebra/DirectSum/Basic.lean | 2 +- Mathlib/Algebra/Group/Subgroup/Ker.lean | 21 ++++++++++++------- .../Algebra/Group/Submonoid/Membership.lean | 2 +- .../Algebra/Group/Submonoid/Operations.lean | 14 +++++++++---- Mathlib/GroupTheory/Congruence/Basic.lean | 2 +- Mathlib/GroupTheory/Finiteness.lean | 4 ++-- Mathlib/GroupTheory/Index.lean | 4 ++-- Mathlib/GroupTheory/PGroup.lean | 2 +- Mathlib/GroupTheory/SchurZassenhaus.lean | 6 +++--- .../SpecificGroups/Alternating.lean | 2 +- Mathlib/GroupTheory/Sylow.lean | 6 +++--- 11 files changed, 38 insertions(+), 27 deletions(-) diff --git a/Mathlib/Algebra/DirectSum/Basic.lean b/Mathlib/Algebra/DirectSum/Basic.lean index 50045a6c1ba5b..83ed3a6e295fa 100644 --- a/Mathlib/Algebra/DirectSum/Basic.lean +++ b/Mathlib/Algebra/DirectSum/Basic.lean @@ -375,7 +375,7 @@ def IsInternal {M S : Type*} [DecidableEq ι] [AddCommMonoid M] [SetLike S M] theorem IsInternal.addSubmonoid_iSup_eq_top {M : Type*} [DecidableEq ι] [AddCommMonoid M] (A : ι → AddSubmonoid M) (h : IsInternal A) : iSup A = ⊤ := by - rw [AddSubmonoid.iSup_eq_mrange_dfinsupp_sumAddHom, AddMonoidHom.mrange_eq_top_iff_surjective] + rw [AddSubmonoid.iSup_eq_mrange_dfinsupp_sumAddHom, AddMonoidHom.mrange_eq_top] exact Function.Bijective.surjective h variable {M S : Type*} [AddCommMonoid M] [SetLike S M] [AddSubmonoidClass S M] diff --git a/Mathlib/Algebra/Group/Subgroup/Ker.lean b/Mathlib/Algebra/Group/Subgroup/Ker.lean index 8b9be58076933..5ba75d983b6b7 100644 --- a/Mathlib/Algebra/Group/Subgroup/Ker.lean +++ b/Mathlib/Algebra/Group/Subgroup/Ker.lean @@ -117,6 +117,9 @@ lemma rangeRestrict_injective_iff {f : G →* N} : Injective f.rangeRestrict ↔ theorem map_range (g : N →* P) (f : G →* N) : f.range.map g = (g.comp f).range := by rw [range_eq_map, range_eq_map]; exact (⊤ : Subgroup G).map_map g f +@[to_additive] +lemma range_comp (g : N →* P) (f : G →* N) : (g.comp f).range = f.range.map g := (map_range ..).symm + @[to_additive] theorem range_eq_top {N} [Group N] {f : G →* N} : f.range = (⊤ : Subgroup N) ↔ Function.Surjective f := @@ -138,10 +141,11 @@ theorem range_one : (1 : G →* N).range = ⊥ := SetLike.ext fun x => by simpa using @comm _ (· = ·) _ 1 x @[to_additive (attr := simp)] -theorem _root_.Subgroup.subtype_range (H : Subgroup G) : H.subtype.range = H := by - rw [range_eq_map, ← SetLike.coe_set_eq, coe_map, Subgroup.coeSubtype] - ext - simp +theorem _root_.Subgroup.range_subtype (H : Subgroup G) : H.subtype.range = H := + SetLike.coe_injective <| (coe_range _).trans <| Subtype.range_coe + +@[to_additive (attr := deprecated (since := "2024-11-26"))] +alias _root_.Subgroup.subtype_range := Subgroup.range_subtype @[to_additive (attr := simp)] theorem _root_.Subgroup.inclusion_range {H K : Subgroup G} (h_le : H ≤ K) : @@ -368,7 +372,7 @@ theorem map_le_range (H : Subgroup G) : map f H ≤ f.range := @[to_additive] theorem map_subtype_le {H : Subgroup G} (K : Subgroup H) : K.map H.subtype ≤ H := - (K.map_le_range H.subtype).trans (le_of_eq H.subtype_range) + (K.map_le_range H.subtype).trans_eq H.range_subtype @[to_additive] theorem ker_le_comap (H : Subgroup N) : f.ker ≤ comap f H := @@ -466,7 +470,7 @@ theorem map_injective_of_ker_le {H K : Subgroup G} (hH : f.ker ≤ H) (hK : f.ke @[to_additive] theorem closure_preimage_eq_top (s : Set G) : closure ((closure s).subtype ⁻¹' s) = ⊤ := by apply map_injective (closure s).subtype_injective - rw [MonoidHom.map_closure, ← MonoidHom.range_eq_map, subtype_range, + rw [MonoidHom.map_closure, ← MonoidHom.range_eq_map, range_subtype, Set.image_preimage_eq_of_subset] rw [coeSubtype, Subtype.range_coe_subtype] exact subset_closure @@ -487,7 +491,8 @@ theorem comap_sup_eq (H K : Subgroup N) (hf : Function.Surjective f) : @[to_additive] theorem sup_subgroupOf_eq {H K L : Subgroup G} (hH : H ≤ L) (hK : K ≤ L) : H.subgroupOf L ⊔ K.subgroupOf L = (H ⊔ K).subgroupOf L := - comap_sup_eq_of_le_range L.subtype (hH.trans L.subtype_range.ge) (hK.trans L.subtype_range.ge) + comap_sup_eq_of_le_range L.subtype (hH.trans_eq L.range_subtype.symm) + (hK.trans_eq L.range_subtype.symm) @[to_additive] theorem codisjoint_subgroupOf_sup (H K : Subgroup G) : @@ -501,7 +506,7 @@ theorem subgroupOf_sup (A A' B : Subgroup G) (hA : A ≤ B) (hA' : A' ≤ B) : refine map_injective_of_ker_le B.subtype (ker_le_comap _ _) (le_trans (ker_le_comap B.subtype _) le_sup_left) ?_ - simp only [subgroupOf, map_comap_eq, map_sup, subtype_range] + simp only [subgroupOf, map_comap_eq, map_sup, range_subtype] rw [inf_of_le_right (sup_le hA hA'), inf_of_le_right hA', inf_of_le_right hA] end Subgroup diff --git a/Mathlib/Algebra/Group/Submonoid/Membership.lean b/Mathlib/Algebra/Group/Submonoid/Membership.lean index b48b123c3e191..2eeb9e1f37c14 100644 --- a/Mathlib/Algebra/Group/Submonoid/Membership.lean +++ b/Mathlib/Algebra/Group/Submonoid/Membership.lean @@ -543,7 +543,7 @@ open MonoidHom @[to_additive] theorem sup_eq_range (s t : Submonoid N) : s ⊔ t = mrange (s.subtype.coprod t.subtype) := by rw [mrange_eq_map, ← mrange_inl_sup_mrange_inr, map_sup, map_mrange, coprod_comp_inl, map_mrange, - coprod_comp_inr, range_subtype, range_subtype] + coprod_comp_inr, mrange_subtype, mrange_subtype] @[to_additive] theorem mem_sup {s t : Submonoid N} {x : N} : x ∈ s ⊔ t ↔ ∃ y ∈ s, ∃ z ∈ t, y * z = x := by diff --git a/Mathlib/Algebra/Group/Submonoid/Operations.lean b/Mathlib/Algebra/Group/Submonoid/Operations.lean index 36ee592c01e9d..86c4066e007df 100644 --- a/Mathlib/Algebra/Group/Submonoid/Operations.lean +++ b/Mathlib/Algebra/Group/Submonoid/Operations.lean @@ -628,6 +628,10 @@ theorem coe_mrange (f : F) : (mrange f : Set N) = Set.range f := theorem mem_mrange {f : F} {y : N} : y ∈ mrange f ↔ ∃ x, f x = y := Iff.rfl +@[to_additive] +lemma mrange_comp {O : Type*} [Monoid O] (f : N →* O) (g : M →* N) : + mrange (f.comp g) = (mrange g).map f := SetLike.coe_injective <| Set.range_comp f _ + @[to_additive] theorem mrange_eq_map (f : F) : mrange f = (⊤ : Submonoid M).map f := Submonoid.copy_eq _ @@ -641,18 +645,18 @@ theorem map_mrange (g : N →* P) (f : M →* N) : f.mrange.map g = mrange (comp simpa only [mrange_eq_map] using (⊤ : Submonoid M).map_map g f @[to_additive] -theorem mrange_eq_top_iff_surjective {f : F} : mrange f = (⊤ : Submonoid N) ↔ Surjective f := +theorem mrange_eq_top {f : F} : mrange f = (⊤ : Submonoid N) ↔ Surjective f := SetLike.ext'_iff.trans <| Iff.trans (by rw [coe_mrange, coe_top]) Set.range_eq_univ @[deprecated (since := "2024-11-11")] -alias mrange_top_iff_surjective := mrange_eq_top_iff_surjective +alias mrange_top_iff_surjective := mrange_eq_top /-- The range of a surjective monoid hom is the whole of the codomain. -/ @[to_additive (attr := simp) "The range of a surjective `AddMonoid` hom is the whole of the codomain."] theorem mrange_eq_top_of_surjective (f : F) (hf : Function.Surjective f) : mrange f = (⊤ : Submonoid N) := - mrange_eq_top_iff_surjective.2 hf + mrange_eq_top.2 hf @[deprecated (since := "2024-11-11")] alias mrange_top_of_surjective := mrange_eq_top_of_surjective @@ -863,9 +867,11 @@ def inclusion {S T : Submonoid M} (h : S ≤ T) : S →* T := S.subtype.codRestrict _ fun x => h x.2 @[to_additive (attr := simp)] -theorem range_subtype (s : Submonoid M) : mrange s.subtype = s := +theorem mrange_subtype (s : Submonoid M) : mrange s.subtype = s := SetLike.coe_injective <| (coe_mrange _).trans <| Subtype.range_coe +@[to_additive (attr := deprecated (since := "2024-11-25"))] alias range_subtype := mrange_subtype + @[to_additive] theorem eq_top_iff' : S = ⊤ ↔ ∀ x : M, x ∈ S := eq_top_iff.trans ⟨fun h m => h <| mem_top m, fun h m _ => h m⟩ diff --git a/Mathlib/GroupTheory/Congruence/Basic.lean b/Mathlib/GroupTheory/Congruence/Basic.lean index 2810c0bd2c101..259eef9b6e583 100644 --- a/Mathlib/GroupTheory/Congruence/Basic.lean +++ b/Mathlib/GroupTheory/Congruence/Basic.lean @@ -173,7 +173,7 @@ variable (x y : M) @[to_additive (attr := simp)] -- Porting note: removed dot notation theorem mrange_mk' : MonoidHom.mrange c.mk' = ⊤ := - MonoidHom.mrange_eq_top_iff_surjective.2 mk'_surjective + MonoidHom.mrange_eq_top.2 mk'_surjective variable {f : M →* P} diff --git a/Mathlib/GroupTheory/Finiteness.lean b/Mathlib/GroupTheory/Finiteness.lean index bd31d6b8743b4..b796773c813c6 100644 --- a/Mathlib/GroupTheory/Finiteness.lean +++ b/Mathlib/GroupTheory/Finiteness.lean @@ -138,7 +138,7 @@ theorem Submonoid.FG.map_injective {M' : Type*} [Monoid M'] {P : Submonoid M} (e @[to_additive (attr := simp)] theorem Monoid.fg_iff_submonoid_fg (N : Submonoid M) : Monoid.FG N ↔ N.FG := by - conv_rhs => rw [← N.range_subtype, MonoidHom.mrange_eq_map] + conv_rhs => rw [← N.mrange_subtype, MonoidHom.mrange_eq_map] exact ⟨fun h => h.out.map N.subtype, fun h => ⟨h.map_injective N.subtype Subtype.coe_injective⟩⟩ @[to_additive] @@ -148,7 +148,7 @@ theorem Monoid.fg_of_surjective {M' : Type*} [Monoid M'] [Monoid.FG M] (f : M obtain ⟨s, hs⟩ := Monoid.fg_def.mp ‹_› use s.image f rwa [Finset.coe_image, ← MonoidHom.map_mclosure, hs, ← MonoidHom.mrange_eq_map, - MonoidHom.mrange_eq_top_iff_surjective] + MonoidHom.mrange_eq_top] @[to_additive] instance Monoid.fg_range {M' : Type*} [Monoid M'] [Monoid.FG M] (f : M →* M') : diff --git a/Mathlib/GroupTheory/Index.lean b/Mathlib/GroupTheory/Index.lean index 0201d5ad30137..95b2a5c3c1451 100644 --- a/Mathlib/GroupTheory/Index.lean +++ b/Mathlib/GroupTheory/Index.lean @@ -80,7 +80,7 @@ theorem index_comap (f : G' →* G) : @[to_additive] theorem relindex_comap (f : G' →* G) (K : Subgroup G') : relindex (comap f H) K = relindex H (map f K) := by - rw [relindex, subgroupOf, comap_comap, index_comap, ← f.map_range, K.subtype_range] + rw [relindex, subgroupOf, comap_comap, index_comap, ← f.map_range, K.range_subtype] variable {H K L} @@ -262,7 +262,7 @@ theorem index_map_of_injective {f : G →* G'} (hf : Function.Injective f) : @[to_additive] theorem index_map_subtype {H : Subgroup G} (K : Subgroup H) : (K.map H.subtype).index = K.index * H.index := by - rw [K.index_map_of_injective H.subtype_injective, H.subtype_range] + rw [K.index_map_of_injective H.subtype_injective, H.range_subtype] @[to_additive] theorem index_eq_card : H.index = Nat.card (G ⧸ H) := diff --git a/Mathlib/GroupTheory/PGroup.lean b/Mathlib/GroupTheory/PGroup.lean index 071291b214ad5..abcab1e3e679d 100644 --- a/Mathlib/GroupTheory/PGroup.lean +++ b/Mathlib/GroupTheory/PGroup.lean @@ -240,7 +240,7 @@ theorem to_inf_right {H K : Subgroup G} (hK : IsPGroup p K) : IsPGroup p (H ⊓ theorem map {H : Subgroup G} (hH : IsPGroup p H) {K : Type*} [Group K] (ϕ : G →* K) : IsPGroup p (H.map ϕ) := by - rw [← H.subtype_range, MonoidHom.map_range] + rw [← H.range_subtype, MonoidHom.map_range] exact hH.of_surjective (ϕ.restrict H).rangeRestrict (ϕ.restrict H).rangeRestrict_surjective theorem comap_of_ker_isPGroup {H : Subgroup G} (hH : IsPGroup p H) {K : Type*} [Group K] diff --git a/Mathlib/GroupTheory/SchurZassenhaus.lean b/Mathlib/GroupTheory/SchurZassenhaus.lean index acdf40f97cdde..e39e21cc46912 100644 --- a/Mathlib/GroupTheory/SchurZassenhaus.lean +++ b/Mathlib/GroupTheory/SchurZassenhaus.lean @@ -174,7 +174,7 @@ private theorem step1 (K : Subgroup G) (hK : K ⊔ N = ⊤) : K = ⊤ := by replace hH : Nat.card (H.map K.subtype) = N.index := by rw [← relindex_bot_left, ← relindex_comap, MonoidHom.comap_bot, Subgroup.ker_subtype, relindex_bot_left, ← IsComplement'.index_eq_card (IsComplement'.symm hH), index_comap, - subtype_range, ← relindex_sup_right, hK, relindex_top_right] + range_subtype, ← relindex_sup_right, hK, relindex_top_right] have h7 : Nat.card N * Nat.card (H.map K.subtype) = Nat.card G := by rw [hH, ← N.index_mul_card, mul_comm] have h8 : (Nat.card N).Coprime (Nat.card (H.map K.subtype)) := by @@ -220,7 +220,7 @@ private theorem step3 (K : Subgroup N) [(K.map N.subtype).Normal] : K = ⊥ ∨ conv at key => rhs rhs - rw [← N.subtype_range, N.subtype.range_eq_map] + rw [← N.range_subtype, N.subtype.range_eq_map] have inj := map_injective N.subtype_injective rwa [inj.eq_iff, inj.eq_iff] at key @@ -239,7 +239,7 @@ include h2 in private theorem step6 : IsPGroup (Nat.card N).minFac N := by haveI : Fact (Nat.card N).minFac.Prime := ⟨step4 h1 h3⟩ refine Sylow.nonempty.elim fun P => P.2.of_surjective P.1.subtype ?_ - rw [← MonoidHom.range_eq_top, subtype_range] + rw [← MonoidHom.range_eq_top, range_subtype] haveI : (P.1.map N.subtype).Normal := normalizer_eq_top.mp (step1 h1 h2 h3 (P.1.map N.subtype).normalizer P.normalizer_sup_eq_top) exact (step3 h1 h2 h3 P.1).resolve_left (step5 h1 h3) diff --git a/Mathlib/GroupTheory/SpecificGroups/Alternating.lean b/Mathlib/GroupTheory/SpecificGroups/Alternating.lean index 8b0f82d484a0d..7dd9feaf90d4d 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Alternating.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Alternating.lean @@ -168,7 +168,7 @@ theorem IsThreeCycle.alternating_normalClosure (h5 : 5 ≤ Fintype.card α) {f : (by have hi : Function.Injective (alternatingGroup α).subtype := Subtype.coe_injective refine eq_top_iff.1 (map_injective hi (le_antisymm (map_mono le_top) ?_)) - rw [← MonoidHom.range_eq_map, subtype_range, normalClosure, MonoidHom.map_closure] + rw [← MonoidHom.range_eq_map, range_subtype, normalClosure, MonoidHom.map_closure] refine (le_of_eq closure_three_cycles_eq_alternating.symm).trans (closure_mono ?_) intro g h obtain ⟨c, rfl⟩ := isConj_iff.1 (isConj_iff_cycleType_eq.2 (hf.trans h.symm)) diff --git a/Mathlib/GroupTheory/Sylow.lean b/Mathlib/GroupTheory/Sylow.lean index fdfa524a631b9..6b4655caeb3b5 100644 --- a/Mathlib/GroupTheory/Sylow.lean +++ b/Mathlib/GroupTheory/Sylow.lean @@ -140,7 +140,7 @@ theorem coe_comapOfInjective (hϕ : Function.Injective ϕ) (h : P ≤ ϕ.range) /-- A sylow subgroup of G is also a sylow subgroup of a subgroup of G. -/ protected def subtype (h : P ≤ N) : Sylow p N := - P.comapOfInjective N.subtype Subtype.coe_injective (by rwa [subtype_range]) + P.comapOfInjective N.subtype Subtype.coe_injective (by rwa [range_subtype]) @[simp] theorem coe_subtype (h : P ≤ N) : P.subtype h = subgroupOf P N := @@ -756,9 +756,9 @@ theorem normalizer_normalizer {p : ℕ} [Fact p.Prime] [Finite (Sylow p G)] (P : have := normal_of_normalizer_normal (P.subtype (le_normalizer.trans le_normalizer)) simp_rw [← normalizer_eq_top, coe_subtype, ← subgroupOf_normalizer_eq le_normalizer, ← subgroupOf_normalizer_eq le_rfl, subgroupOf_self] at this - rw [← subtype_range P.normalizer.normalizer, MonoidHom.range_eq_map, + rw [← range_subtype P.normalizer.normalizer, MonoidHom.range_eq_map, ← this trivial] - exact map_comap_eq_self (le_normalizer.trans (ge_of_eq (subtype_range _))) + exact map_comap_eq_self (le_normalizer.trans (ge_of_eq (range_subtype _))) theorem normal_of_all_max_subgroups_normal [Finite G] (hnc : ∀ H : Subgroup G, IsCoatom H → H.Normal) {p : ℕ} [Fact p.Prime] [Finite (Sylow p G)] From cb86c7503d19e5d7a123067aec02599a8eaad66d Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Wed, 27 Nov 2024 11:40:18 +0000 Subject: [PATCH 351/829] feat(CategoryTheory): pushforward pullback adjunction for `P.Over Q X` (#19271) --- Mathlib.lean | 1 + .../Limits/MorphismProperty.lean | 3 +- .../MorphismProperty/Basic.lean | 4 + .../MorphismProperty/Comma.lean | 101 ++++++++++++- .../MorphismProperty/OverAdjunction.lean | 133 ++++++++++++++++++ 5 files changed, 238 insertions(+), 4 deletions(-) create mode 100644 Mathlib/CategoryTheory/MorphismProperty/OverAdjunction.lean diff --git a/Mathlib.lean b/Mathlib.lean index f3e568ce173a2..a843c7987f61d 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1991,6 +1991,7 @@ import Mathlib.CategoryTheory.MorphismProperty.Concrete import Mathlib.CategoryTheory.MorphismProperty.Factorization import Mathlib.CategoryTheory.MorphismProperty.IsInvertedBy import Mathlib.CategoryTheory.MorphismProperty.Limits +import Mathlib.CategoryTheory.MorphismProperty.OverAdjunction import Mathlib.CategoryTheory.MorphismProperty.Representable import Mathlib.CategoryTheory.NatIso import Mathlib.CategoryTheory.NatTrans diff --git a/Mathlib/CategoryTheory/Limits/MorphismProperty.lean b/Mathlib/CategoryTheory/Limits/MorphismProperty.lean index b4bb842c72dd2..48a845bacb238 100644 --- a/Mathlib/CategoryTheory/Limits/MorphismProperty.lean +++ b/Mathlib/CategoryTheory/Limits/MorphismProperty.lean @@ -119,9 +119,8 @@ instance [P.ContainsIdentities] (Y : P.Over ⊤ X) : uniq a := by ext · simp only [mk_left, Hom.hom_left, homMk_hom, Over.homMk_left] - rw [← Over.w a.hom] + rw [← Over.w a] simp only [mk_left, Functor.const_obj_obj, Hom.hom_left, mk_hom, Category.comp_id] - · rfl /-- `X ⟶ X` is the terminal object of `P.Over ⊤ X`. -/ def mkIdTerminal [P.ContainsIdentities] : diff --git a/Mathlib/CategoryTheory/MorphismProperty/Basic.lean b/Mathlib/CategoryTheory/MorphismProperty/Basic.lean index cf9f01545d900..e8a96be14d9ab 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/Basic.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/Basic.lean @@ -157,6 +157,10 @@ lemma RespectsIso.precomp (P : MorphismProperty C) [P.RespectsIso] {X Y Z : C} ( [IsIso e] (f : Y ⟶ Z) (hf : P f) : P (e ≫ f) := RespectsLeft.precomp (Q := isomorphisms C) e ‹IsIso e› f hf +instance : RespectsIso (⊤ : MorphismProperty C) where + precomp _ _ _ _ := trivial + postcomp _ _ _ _ := trivial + lemma RespectsIso.postcomp (P : MorphismProperty C) [P.RespectsIso] {X Y Z : C} (e : Y ⟶ Z) [IsIso e] (f : X ⟶ Y) (hf : P f) : P (f ≫ e) := RespectsRight.postcomp (Q := isomorphisms C) e ‹IsIso e› f hf diff --git a/Mathlib/CategoryTheory/MorphismProperty/Comma.lean b/Mathlib/CategoryTheory/MorphismProperty/Comma.lean index 6b093565af86f..3dd55847c9890 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/Comma.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/Comma.lean @@ -79,10 +79,8 @@ lemma Hom.hom_mk {X Y : P.Comma L R Q W} (f : CommaMorphism X.toComma Y.toComma) (hf) (hg) : Comma.Hom.hom ⟨f, hf, hg⟩ = f := rfl -@[simp] lemma Hom.hom_left {X Y : P.Comma L R Q W} (f : Comma.Hom X Y) : f.hom.left = f.left := rfl -@[simp] lemma Hom.hom_right {X Y : P.Comma L R Q W} (f : Comma.Hom X Y) : f.hom.right = f.right := rfl /-- See Note [custom simps projection] -/ @@ -133,6 +131,14 @@ lemma id_hom (X : P.Comma L R Q W) : (𝟙 X : X ⟶ X).hom = 𝟙 X.toComma := lemma comp_hom {X Y Z : P.Comma L R Q W} (f : X ⟶ Y) (g : Y ⟶ Z) : (f ≫ g).hom = f.hom ≫ g.hom := rfl +@[reassoc] +lemma comp_left {X Y Z : P.Comma L R Q W} (f : X ⟶ Y) (g : Y ⟶ Z) : + (f ≫ g).left = f.left ≫ g.left := rfl + +@[reassoc] +lemma comp_right {X Y Z : P.Comma L R Q W} (f : X ⟶ Y) (g : Y ⟶ Z) : + (f ≫ g).right = f.right ≫ g.right := rfl + /-- If `i` is an isomorphism in `Comma L R`, it is also a morphism in `P.Comma L R Q W`. -/ @[simps hom] def homFromCommaOfIsIso [Q.RespectsIso] [W.RespectsIso] {X Y : P.Comma L R Q W} @@ -205,6 +211,63 @@ def forgetFullyFaithful : (forget L R P ⊤ ⊤).FullyFaithful where instance : (forget L R P ⊤ ⊤).Full := Functor.FullyFaithful.full (forgetFullyFaithful L R P) +section + +variable {L R} + +@[simp] +lemma eqToHom_left {X Y : P.Comma L R Q W} (h : X = Y) : + (eqToHom h).left = eqToHom (by rw [h]) := by + subst h + rfl + +@[simp] +lemma eqToHom_right {X Y : P.Comma L R Q W} (h : X = Y) : + (eqToHom h).right = eqToHom (by rw [h]) := by + subst h + rfl + +end + +section Functoriality + +variable {L R P Q W} +variable {L₁ L₂ L₃ : A ⥤ T} {R₁ R₂ R₃ : B ⥤ T} + +/-- Lift a functor `F : C ⥤ Comma L R` to the subcategory `P.Comma L R Q W` under +suitable assumptions on `F`. -/ +@[simps obj_toComma map_hom] +def lift {C : Type*} [Category C] (F : C ⥤ Comma L R) + (hP : ∀ X, P (F.obj X).hom) + (hQ : ∀ {X Y} (f : X ⟶ Y), Q (F.map f).left) + (hW : ∀ {X Y} (f : X ⟶ Y), W (F.map f).right) : + C ⥤ P.Comma L R Q W where + obj X := + { __ := F.obj X + prop := hP X } + map {X Y} f := + { __ := F.map f + prop_hom_left := hQ f + prop_hom_right := hW f } + +variable (R) in +/-- A natural transformation `L₁ ⟶ L₂` induces a functor `P.Comma L₂ R Q W ⥤ P.Comma L₁ R Q W`. -/ +@[simps!] +def mapLeft (l : L₁ ⟶ L₂) (hl : ∀ X : P.Comma L₂ R Q W, P (l.app X.left ≫ X.hom)) : + P.Comma L₂ R Q W ⥤ P.Comma L₁ R Q W := + lift (forget _ _ _ _ _ ⋙ CategoryTheory.Comma.mapLeft R l) hl + (fun f ↦ f.prop_hom_left) (fun f ↦ f.prop_hom_right) + +variable (L) in +/-- A natural transformation `R₁ ⟶ R₂` induces a functor `P.Comma L R₁ Q W ⥤ P.Comma L R₂ Q W`. -/ +@[simps!] +def mapRight (r : R₁ ⟶ R₂) (hr : ∀ X : P.Comma L R₁ Q W, P (X.hom ≫ r.app X.right)) : + P.Comma L R₁ Q W ⥤ P.Comma L R₂ Q W := + lift (forget _ _ _ _ _ ⋙ CategoryTheory.Comma.mapRight L r) hr + (fun f ↦ f.prop_hom_left) (fun f ↦ f.prop_hom_right) + +end Functoriality + end Comma end Comma @@ -251,6 +314,23 @@ protected def Over.homMk {A B : P.Over Q X} (f : A.left ⟶ B.left) prop_hom_left := hf prop_hom_right := trivial +/-- Make an isomorphism in `P.Over Q X` from an isomorphism in `T` with compatibilities. -/ +@[simps! hom_left inv_left] +protected def Over.isoMk [Q.RespectsIso] {A B : P.Over Q X} (f : A.left ≅ B.left) + (w : f.hom ≫ B.hom = A.hom := by aesop_cat) : A ≅ B := + Comma.isoMk f (Discrete.eqToIso' rfl) + +@[ext] +lemma Over.Hom.ext {A B : P.Over Q X} {f g : A ⟶ B} (h : f.left = g.left) : f = g := by + ext + · exact h + · simp + +@[reassoc] +lemma Over.w {A B : P.Over Q X} (f : A ⟶ B) : + f.left ≫ B.hom = A.hom := by + simp + end Over section Under @@ -295,6 +375,23 @@ protected def Under.homMk {A B : P.Under Q X} (f : A.right ⟶ B.right) prop_hom_left := trivial prop_hom_right := hf +/-- Make an isomorphism in `P.Under Q X` from an isomorphism in `T` with compatibilities. -/ +@[simps! hom_right inv_right] +protected def Under.isoMk [Q.RespectsIso] {A B : P.Under Q X} (f : A.right ≅ B.right) + (w : A.hom ≫ f.hom = B.hom := by aesop_cat) : A ≅ B := + Comma.isoMk (Discrete.eqToIso' rfl) f + +@[ext] +lemma Under.Hom.ext {A B : P.Under Q X} {f g : A ⟶ B} (h : f.right = g.right) : f = g := by + ext + · simp + · exact h + +@[reassoc] +lemma Under.w {A B : P.Under Q X} (f : A ⟶ B) : + A.hom ≫ f.right = B.hom := by + simp + end Under end CategoryTheory.MorphismProperty diff --git a/Mathlib/CategoryTheory/MorphismProperty/OverAdjunction.lean b/Mathlib/CategoryTheory/MorphismProperty/OverAdjunction.lean new file mode 100644 index 0000000000000..4a41ebda2533c --- /dev/null +++ b/Mathlib/CategoryTheory/MorphismProperty/OverAdjunction.lean @@ -0,0 +1,133 @@ +/- +Copyright (c) 2024 Christian Merten. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Christian Merten +-/ +import Mathlib.CategoryTheory.MorphismProperty.Comma +import Mathlib.CategoryTheory.Adjunction.Over +import Mathlib.CategoryTheory.MorphismProperty.Limits + +/-! +# Adjunction of pushforward and pullback in `P.Over Q X` + +A morphism `f : X ⟶ Y` defines two functors: + +- `Over.map`: post-composition with `f` +- `Over.pullback`: base-change along `f` + +These are adjoint under suitable assumptions on `P` and `Q`. + +-/ + +namespace CategoryTheory.MorphismProperty + +open Limits + +variable {T : Type*} [Category T] (P Q : MorphismProperty T) [Q.IsMultiplicative] +variable {X Y : T} (f : X ⟶ Y) + +section Map + +variable {P} [P.IsStableUnderComposition] (hPf : P f) + +variable {f} + +/-- If `P` is stable under composition and `f : X ⟶ Y` satisfies `P`, +this is the functor `P.Over Q X ⥤ P.Over Q Y` given by composing with `f`. -/ +@[simps! obj_left obj_hom map_left] +def Over.map : P.Over Q X ⥤ P.Over Q Y := + Comma.mapRight _ (Discrete.natTrans fun _ ↦ f) <| fun X ↦ P.comp_mem _ _ X.prop hPf + +lemma Over.map_comp {X Y Z : T} {f : X ⟶ Y} (hf : P f) {g : Y ⟶ Z} (hg : P g) : + map Q (P.comp_mem f g hf hg) = map Q hf ⋙ map Q hg := by + fapply Functor.ext + · simp [map, Comma.mapRight, CategoryTheory.Comma.mapRight, Comma.lift] + · intro U V k + ext + simp + +/-- `Over.map` commutes with composition. -/ +@[simps! hom_app_left inv_app_left] +def Over.mapComp {X Y Z : T} {f : X ⟶ Y} (hf : P f) {g : Y ⟶ Z} (hg : P g) [Q.RespectsIso] : + map Q (P.comp_mem f g hf hg) ≅ map Q hf ⋙ map Q hg := + NatIso.ofComponents (fun X ↦ Over.isoMk (Iso.refl _)) + +end Map + +section Pullback + +variable [HasPullbacks T] [P.IsStableUnderBaseChange] [Q.IsStableUnderBaseChange] + +/-- If `P` and `Q` are stable under base change and pullbacks exist in `T`, +this is the functor `P.Over Q Y ⥤ P.Over Q X` given by base change along `f`. -/ +@[simps! obj_left obj_hom map_left] +noncomputable def Over.pullback : P.Over Q Y ⥤ P.Over Q X where + obj A := + { __ := (CategoryTheory.Over.pullback f).obj A.toComma + prop := P.pullback_snd _ _ A.prop } + map {A B} g := + { __ := (CategoryTheory.Over.pullback f).map g.toCommaMorphism + prop_hom_left := Q.baseChange_map f g.toCommaMorphism g.prop_hom_left + prop_hom_right := trivial } + +variable {P} {Q} + +/-- `Over.pullback` commutes with composition. -/ +@[simps! hom_app_left inv_app_left] +noncomputable def Over.pullbackComp [Q.RespectsIso] {X Y Z : T} (f : X ⟶ Y) (g : Y ⟶ Z) : + Over.pullback P Q (f ≫ g) ≅ Over.pullback P Q g ⋙ Over.pullback P Q f := + NatIso.ofComponents + (fun X ↦ Over.isoMk ((pullbackLeftPullbackSndIso X.hom g f).symm) (by simp)) + +lemma Over.pullbackComp_left_fst_fst [Q.RespectsIso] {X Y Z : T} (f : X ⟶ Y) (g : Y ⟶ Z) + (A : P.Over Q Z) : + ((Over.pullbackComp f g).hom.app A).left ≫ + pullback.fst (pullback.snd A.hom g) f ≫ pullback.fst A.hom g = + pullback.fst A.hom (f ≫ g) := by + simp + +/-- If `f = g`, then base change along `f` is naturally isomorphic to base change along `g`. -/ +noncomputable def Over.pullbackCongr {X Y : T} {f g : X ⟶ Y} (h : f = g) : + Over.pullback P Q f ≅ Over.pullback P Q g := + NatIso.ofComponents (fun X ↦ eqToIso (by rw [h])) + +@[reassoc (attr := simp)] +lemma Over.pullbackCongr_hom_app_left_fst {X Y : T} {f g : X ⟶ Y} (h : f = g) (A : P.Over Q Y) : + ((Over.pullbackCongr h).hom.app A).left ≫ pullback.fst A.hom g = + pullback.fst A.hom f := by + subst h + simp [pullbackCongr] + +end Pullback + +section Adjunction + +variable [P.IsStableUnderComposition] [P.IsStableUnderBaseChange] + [Q.IsStableUnderBaseChange] [HasPullbacks T] + +/-- `P.Over.map` is left adjoint to `P.Over.pullback` if `f` satisfies `P`. -/ +noncomputable def Over.mapPullbackAdj [Q.HasOfPostcompProperty Q] (hPf : P f) (hQf : Q f) : + Over.map Q hPf ⊣ Over.pullback P Q f := + Adjunction.mkOfHomEquiv + { homEquiv := fun A B ↦ + { toFun := fun g ↦ + Over.homMk (pullback.lift g.left A.hom <| by simp) (by simp) <| by + apply Q.of_postcomp (W' := Q) + · exact Q.pullback_fst B.hom f hQf + · simpa using g.prop_hom_left + invFun := fun h ↦ Over.homMk (h.left ≫ pullback.fst B.hom f) + (by + simp only [map_obj_left, Functor.const_obj_obj, pullback_obj_left, Functor.id_obj, + Category.assoc, pullback.condition, map_obj_hom, ← pullback_obj_hom, Over.w_assoc]) + (Q.comp_mem _ _ h.prop_hom_left (Q.pullback_fst _ _ hQf)) + left_inv := by aesop_cat + right_inv := fun h ↦ by + ext + dsimp + ext + · simp + · simpa using h.w.symm } } + +end Adjunction + +end CategoryTheory.MorphismProperty From 6b724bf9d816ee5143e5b8b8e8be69b40b820656 Mon Sep 17 00:00:00 2001 From: blizzard_inc Date: Wed, 27 Nov 2024 11:40:20 +0000 Subject: [PATCH 352/829] fix(scripts/zulip_emoji_merge_delegate.py): fix the emoji reaction script (#19535) change the script to use `True` rather than `true`, fixing an error. --- scripts/zulip_emoji_merge_delegate.py | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/scripts/zulip_emoji_merge_delegate.py b/scripts/zulip_emoji_merge_delegate.py index cbecbdf203c10..91836a41e49ff 100644 --- a/scripts/zulip_emoji_merge_delegate.py +++ b/scripts/zulip_emoji_merge_delegate.py @@ -30,7 +30,7 @@ "num_before": 5000, "num_after": 0, "narrow": [ - {"operator": "channel", "operand": "rss", "negated": true}, + {"operator": "channel", "operand": "rss", "negated": True}, {"operator": "search", "operand": f'#{PR_NUMBER}'}, ], }) From 671c09054c3b746a0d22ca24711dac1068eb282a Mon Sep 17 00:00:00 2001 From: grunweg Date: Wed, 27 Nov 2024 12:13:29 +0000 Subject: [PATCH 353/829] chore: move style linters to `Style` (#19529) This makes for a more coherent split of the linter files. Third attempt after #15942 and #16510. --- Mathlib/Tactic/Linter/Lint.lean | 337 +---------------------------- Mathlib/Tactic/Linter/Style.lean | 355 ++++++++++++++++++++++++++++++- MathlibTest/Lint.lean | 221 ------------------- MathlibTest/LintStyle.lean | 218 ++++++++++++++++++- MathlibTest/LongFile.lean | 2 +- 5 files changed, 573 insertions(+), 560 deletions(-) diff --git a/Mathlib/Tactic/Linter/Lint.lean b/Mathlib/Tactic/Linter/Lint.lean index 280dad9777e59..67ccae0115ab1 100644 --- a/Mathlib/Tactic/Linter/Lint.lean +++ b/Mathlib/Tactic/Linter/Lint.lean @@ -9,7 +9,8 @@ import Mathlib.Tactic.DeclarationNames /-! # Linters for Mathlib -In this file we define additional linters for mathlib. +In this file we define additional linters for mathlib, +which concern the *behaviour* of the linted code, and not issues of code style or formatting. Perhaps these should be moved to Batteries in the future. -/ @@ -95,336 +96,4 @@ def dupNamespace : Linter where run := withSetOptionIn fun stx ↦ do initialize addLinter dupNamespace -end DupNamespaceLinter - -/-! -# The "missing end" linter - -The "missing end" linter emits a warning on non-closed `section`s and `namespace`s. -It allows the "outermost" `noncomputable section` to be left open (whether or not it is named). --/ - -open Lean Elab Command - -/-- The "missing end" linter emits a warning on non-closed `section`s and `namespace`s. -It allows the "outermost" `noncomputable section` to be left open (whether or not it is named). --/ -register_option linter.style.missingEnd : Bool := { - defValue := false - descr := "enable the missing end linter" -} - -namespace Style.missingEnd - -@[inherit_doc Mathlib.Linter.linter.style.missingEnd] -def missingEndLinter : Linter where run := withSetOptionIn fun stx ↦ do - -- Only run this linter at the end of a module. - unless stx.isOfKind ``Lean.Parser.Command.eoi do return - if Linter.getLinterValue linter.style.missingEnd (← getOptions) && - !(← MonadState.get).messages.hasErrors then - let sc ← getScopes - -- The last scope is always the "base scope", corresponding to no active `section`s or - -- `namespace`s. We are interested in any *other* unclosed scopes. - if sc.length == 1 then return - let ends := sc.dropLast.map fun s ↦ (s.header, s.isNoncomputable) - -- If the outermost scope corresponds to a `noncomputable section`, we ignore it. - let ends := if ends.getLast!.2 then ends.dropLast else ends - -- If there are any further un-closed scopes, we emit a warning. - if !ends.isEmpty then - let ending := (ends.map Prod.fst).foldl (init := "") fun a b ↦ - a ++ s!"\n\nend{if b == "" then "" else " "}{b}" - Linter.logLint linter.style.missingEnd stx - m!"unclosed sections or namespaces; expected: '{ending}'" - -initialize addLinter missingEndLinter - -end Style.missingEnd - -/-! -# The `cdot` linter - -The `cdot` linter is a syntax-linter that flags uses of the "cdot" `·` that are achieved -by typing a character different from `·`. -For instance, a "plain" dot `.` is allowed syntax, but is flagged by the linter. -It also flags "isolated cdots", i.e. when the `·` is on its own line. --/ - -/-- -The `cdot` linter flags uses of the "cdot" `·` that are achieved by typing a character -different from `·`. -For instance, a "plain" dot `.` is allowed syntax, but is flagged by the linter. -It also flags "isolated cdots", i.e. when the `·` is on its own line. -/ -register_option linter.style.cdot : Bool := { - defValue := false - descr := "enable the `cdot` linter" -} - -/-- `isCDot? stx` checks whether `stx` is a `Syntax` node corresponding to a `cdot` typed with -the character `·`. -/ -def isCDot? : Syntax → Bool - | .node _ ``cdotTk #[.node _ `patternIgnore #[.node _ _ #[.atom _ v]]] => v == "·" - | .node _ ``Lean.Parser.Term.cdot #[.atom _ v] => v == "·" - | _ => false - -/-- -`findCDot stx` extracts from `stx` the syntax nodes of `kind` `Lean.Parser.Term.cdot` or `cdotTk`. -/ -partial -def findCDot : Syntax → Array Syntax - | stx@(.node _ kind args) => - let dargs := (args.map findCDot).flatten - match kind with - | ``Lean.Parser.Term.cdot | ``cdotTk => dargs.push stx - | _ => dargs - |_ => #[] - -/-- `unwanted_cdot stx` returns an array of syntax atoms within `stx` -corresponding to `cdot`s that are not written with the character `·`. -This is precisely what the `cdot` linter flags. --/ -def unwanted_cdot (stx : Syntax) : Array Syntax := - (findCDot stx).filter (!isCDot? ·) - -namespace Style - -@[inherit_doc linter.style.cdot] -def cdotLinter : Linter where run := withSetOptionIn fun stx ↦ do - unless Linter.getLinterValue linter.style.cdot (← getOptions) do - return - if (← MonadState.get).messages.hasErrors then - return - for s in unwanted_cdot stx do - Linter.logLint linter.style.cdot s - m!"Please, use '·' (typed as `\\.`) instead of '{s}' as 'cdot'." - -- We also check for isolated cdot's, i.e. when the cdot is on its own line. - for cdot in Mathlib.Linter.findCDot stx do - match cdot.find? (·.isOfKind `token.«· ») with - | some (.node _ _ #[.atom (.original _ _ afterCDot _) _]) => - if (afterCDot.takeWhile (·.isWhitespace)).contains '\n' then - logWarningAt cdot <| .tagged linter.style.cdot.name - m!"This central dot `·` is isolated; please merge it with the next line." - | _ => return - -initialize addLinter cdotLinter - -end Style - -/-! -# The `dollarSyntax` linter - -The `dollarSyntax` linter flags uses of `<|` that are achieved by typing `$`. -These are disallowed by the mathlib style guide, as using `<|` pairs better with `|>`. --/ - -/-- The `dollarSyntax` linter flags uses of `<|` that are achieved by typing `$`. -These are disallowed by the mathlib style guide, as using `<|` pairs better with `|>`. -/ -register_option linter.style.dollarSyntax : Bool := { - defValue := false - descr := "enable the `dollarSyntax` linter" -} - -namespace Style.dollarSyntax - -/-- `findDollarSyntax stx` extracts from `stx` the syntax nodes of `kind` `$`. -/ -partial -def findDollarSyntax : Syntax → Array Syntax - | stx@(.node _ kind args) => - let dargs := (args.map findDollarSyntax).flatten - match kind with - | ``«term_$__» => dargs.push stx - | _ => dargs - |_ => #[] - -@[inherit_doc linter.style.dollarSyntax] -def dollarSyntaxLinter : Linter where run := withSetOptionIn fun stx ↦ do - unless Linter.getLinterValue linter.style.dollarSyntax (← getOptions) do - return - if (← MonadState.get).messages.hasErrors then - return - for s in findDollarSyntax stx do - Linter.logLint linter.style.dollarSyntax s - m!"Please use '<|' instead of '$' for the pipe operator." - -initialize addLinter dollarSyntaxLinter - -end Style.dollarSyntax - -/-! -# The `lambdaSyntax` linter - -The `lambdaSyntax` linter is a syntax linter that flags uses of the symbol `λ` to define anonymous -functions, as opposed to the `fun` keyword. These are syntactically equivalent; mathlib style -prefers the latter as it is considered more readable. --/ - -/-- -The `lambdaSyntax` linter flags uses of the symbol `λ` to define anonymous functions. -This is syntactically equivalent to the `fun` keyword; mathlib style prefers using the latter. --/ -register_option linter.style.lambdaSyntax : Bool := { - defValue := false - descr := "enable the `lambdaSyntax` linter" -} - -namespace Style.lambdaSyntax - -/-- -`findLambdaSyntax stx` extracts from `stx` all syntax nodes of `kind` `Term.fun`. -/ -partial -def findLambdaSyntax : Syntax → Array Syntax - | stx@(.node _ kind args) => - let dargs := (args.map findLambdaSyntax).flatten - match kind with - | ``Parser.Term.fun => dargs.push stx - | _ => dargs - |_ => #[] - -@[inherit_doc linter.style.lambdaSyntax] -def lambdaSyntaxLinter : Linter where run := withSetOptionIn fun stx ↦ do - unless Linter.getLinterValue linter.style.lambdaSyntax (← getOptions) do - return - if (← MonadState.get).messages.hasErrors then - return - for s in findLambdaSyntax stx do - if let .atom _ "λ" := s[0] then - Linter.logLint linter.style.lambdaSyntax s[0] m!"\ - Please use 'fun' and not 'λ' to define anonymous functions.\n\ - The 'λ' syntax is deprecated in mathlib4." - -initialize addLinter lambdaSyntaxLinter - -end Style.lambdaSyntax - -/-! -# The "longFile" linter - -The "longFile" linter emits a warning on files which are longer than a certain number of lines -(1500 by default). --/ - -/-- -The "longFile" linter emits a warning on files which are longer than a certain number of lines -(`linter.style.longFileDefValue` by default on mathlib, no limit for downstream projects). -If this option is set to `N` lines, the linter warns once a file has more than `N` lines. -A value of `0` silences the linter entirely. --/ -register_option linter.style.longFile : Nat := { - defValue := 0 - 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] -def longFileLinter : Linter where run := withSetOptionIn fun stx ↦ do - let linterBound := linter.style.longFile.get (← getOptions) - if linterBound == 0 then - return - let defValue := linter.style.longFileDefValue.get (← getOptions) - let smallOption := match stx with - | `(set_option linter.style.longFile $x) => TSyntax.getNat ⟨x.raw⟩ ≤ defValue - | _ => false - if smallOption then - logWarningAt stx <| .tagged linter.style.longFile.name - m!"The default value of the `longFile` linter is {defValue}.\n\ - The current value of {linterBound} does not exceed the allowed bound.\n\ - Please, remove the `set_option linter.style.longFile {linterBound}`." - else - -- Thanks to the above check, the linter option is either not set (and hence equal - -- to the default) or set to some value *larger* than the default. - -- `Parser.isTerminalCommand` allows `stx` to be `#exit`: this is useful for tests. - unless Parser.isTerminalCommand stx do return - -- We exclude `Mathlib.lean` from the linter: it exceeds linter's default number of allowed - -- lines, and it is an auto-generated import-only file. - -- TODO: if there are more such files, revise the implementation. - if (← getMainModule) == `Mathlib then return - if let some init := stx.getTailPos? then - -- the last line: we subtract 1, since the last line is expected to be empty - let lastLine := ((← getFileMap).toPosition init).line - -- In this case, the file has an allowed length, and the linter option is unnecessarily set. - if lastLine ≤ defValue && defValue < linterBound then - logWarningAt stx <| .tagged linter.style.longFile.name - m!"The default value of the `longFile` linter is {defValue}.\n\ - This file is {lastLine} lines long which does not exceed the allowed bound.\n\ - Please, remove the `set_option linter.style.longFile {linterBound}`." - else - -- `candidate` is divisible by `100` and satisfies `lastLine + 100 < candidate ≤ lastLine + 200` - -- note that either `lastLine ≤ defValue` and `defValue = linterBound` hold or - -- `candidate` is necessarily bigger than `lastLine` and hence bigger than `defValue` - let candidate := (lastLine / 100) * 100 + 200 - let candidate := max candidate defValue - -- In this case, the file is longer than the default and also than what the option says. - if defValue ≤ linterBound && linterBound < lastLine then - logWarningAt stx <| .tagged linter.style.longFile.name - m!"This file is {lastLine} lines long, but the limit is {linterBound}.\n\n\ - You can extend the allowed length of the file using \ - `set_option linter.style.longFile {candidate}`.\n\ - You can completely disable this linter by setting the length limit to `0`." - else - -- Finally, the file exceeds the default value, but not the option: we only allow the value - -- of the option to be `candidate` or `candidate + 100`. - -- In particular, this flags any option that is set to an unnecessarily high value. - if linterBound == candidate || linterBound + 100 == candidate then return - else - logWarningAt stx <| .tagged linter.style.longFile.name - m!"This file is {lastLine} lines long. \ - The current limit is {linterBound}, but it is expected to be {candidate}:\n\ - `set_option linter.style.longFile {candidate}`." - -initialize addLinter longFileLinter - -end Style.longFile - -/-! # The "longLine linter" -/ - -/-- The "longLine" linter emits a warning on lines longer than 100 characters. -We allow lines containing URLs to be longer, though. -/ -register_option linter.style.longLine : Bool := { - defValue := false - descr := "enable the longLine linter" -} - -namespace Style.longLine - -@[inherit_doc Mathlib.Linter.linter.style.longLine] -def longLineLinter : Linter where run := withSetOptionIn fun stx ↦ do - unless Linter.getLinterValue linter.style.longLine (← getOptions) do - return - if (← MonadState.get).messages.hasErrors then - return - -- The linter ignores the `#guard_msgs` command, in particular its doc-string. - -- The linter still lints the message guarded by `#guard_msgs`. - if stx.isOfKind ``Lean.guardMsgsCmd then - return - -- if the linter reached the end of the file, then we scan the `import` syntax instead - let stx := ← do - if stx.isOfKind ``Lean.Parser.Command.eoi then - let fileMap ← getFileMap - -- `impMods` is the syntax for the modules imported in the current file - let (impMods, _) ← Parser.parseHeader - { input := fileMap.source, fileName := ← getFileName, fileMap := fileMap } - return impMods - else return stx - let sstr := stx.getSubstring? - let fm ← getFileMap - let longLines := ((sstr.getD default).splitOn "\n").filter fun line ↦ - (100 < (fm.toPosition line.stopPos).column) - for line in longLines do - if (line.splitOn "http").length ≤ 1 then - let stringMsg := if line.contains '"' then - "\nYou can use \"string gaps\" to format long strings: within a string quotation, \ - using a '\\' at the end of a line allows you to continue the string on the following \ - line, removing all intervening whitespace." - else "" - Linter.logLint linter.style.longLine (.ofRange ⟨line.startPos, line.stopPos⟩) - m!"This line exceeds the 100 character limit, please shorten it!{stringMsg}" -initialize addLinter longLineLinter - -end Style.longLine - -end Mathlib.Linter +end Mathlib.Linter.DupNamespaceLinter diff --git a/Mathlib/Tactic/Linter/Style.lean b/Mathlib/Tactic/Linter/Style.lean index 112d95c09ea65..65272ba2d6d0a 100644 --- a/Mathlib/Tactic/Linter/Style.lean +++ b/Mathlib/Tactic/Linter/Style.lean @@ -12,10 +12,29 @@ import Mathlib.Tactic.Linter.Header /-! ## Style linters -This file contains (currently one, eventually more) linters about stylistic aspects: -these are only about coding style, but do not affect correctness nor global coherence of mathlib. +This file contain linters about stylistic aspects: these are only about coding style, +but do not affect correctness nor global coherence of mathlib. +Historically, some of these were ported from the `lint-style.py` Python script. -Historically, these were ported from the `lint-style.py` Python script. +This file defines the following linters: +- the `set_option` linter checks for the presence of `set_option` commands activating +options disallowed in mathlib: these are meant to be temporary, and not for polished code +- the `missingEnd` linter checks for sections or namespaces which are not closed by the end +of the file: enforcing this invariant makes minimising files or moving code between files easier +- the `cdotLinter` linter checks for focusing dots `·` which are typed using a `.` instead: +this is allowed Lean syntax, but it is nicer to be uniform +- the `dollarSyntax` linter checks for use of the dollar sign `$` instead of the `<|` pipe operator: +similarly, both symbols have the same meaning, but mathlib prefers `<|` for the symmetry with +the `|>` symbol +- the `lambdaSyntax` linter checks for uses of the `λ` symbol for ananomous functions, +instead of the `fun` keyword: mathlib prefers the latter for reasons of readability +(This linter is still under review in PR #15896.) +- the `longLine` linter checks for lines which have more than 100 characters +- the `longFile` linter checks for files which have more than 1500 lines: +this linter is still under development in PR #15610. + +All of these linters are enabled in mathlib by default, but disabled globally +since they enforce conventions which are inherently subjective. -/ open Lean Elab Command @@ -71,4 +90,334 @@ initialize addLinter setOptionLinter end Style.setOption +/-! +# The "missing end" linter + +The "missing end" linter emits a warning on non-closed `section`s and `namespace`s. +It allows the "outermost" `noncomputable section` to be left open (whether or not it is named). +-/ + +open Lean Elab Command + +/-- The "missing end" linter emits a warning on non-closed `section`s and `namespace`s. +It allows the "outermost" `noncomputable section` to be left open (whether or not it is named). +-/ +register_option linter.style.missingEnd : Bool := { + defValue := false + descr := "enable the missing end linter" +} + +namespace Style.missingEnd + +@[inherit_doc Mathlib.Linter.linter.style.missingEnd] +def missingEndLinter : Linter where run := withSetOptionIn fun stx ↦ do + -- Only run this linter at the end of a module. + unless stx.isOfKind ``Lean.Parser.Command.eoi do return + if Linter.getLinterValue linter.style.missingEnd (← getOptions) && + !(← MonadState.get).messages.hasErrors then + let sc ← getScopes + -- The last scope is always the "base scope", corresponding to no active `section`s or + -- `namespace`s. We are interested in any *other* unclosed scopes. + if sc.length == 1 then return + let ends := sc.dropLast.map fun s ↦ (s.header, s.isNoncomputable) + -- If the outermost scope corresponds to a `noncomputable section`, we ignore it. + let ends := if ends.getLast!.2 then ends.dropLast else ends + -- If there are any further un-closed scopes, we emit a warning. + if !ends.isEmpty then + let ending := (ends.map Prod.fst).foldl (init := "") fun a b ↦ + a ++ s!"\n\nend{if b == "" then "" else " "}{b}" + Linter.logLint linter.style.missingEnd stx + m!"unclosed sections or namespaces; expected: '{ending}'" + +initialize addLinter missingEndLinter + +end Style.missingEnd + +/-! +# The `cdot` linter + +The `cdot` linter is a syntax-linter that flags uses of the "cdot" `·` that are achieved +by typing a character different from `·`. +For instance, a "plain" dot `.` is allowed syntax, but is flagged by the linter. +It also flags "isolated cdots", i.e. when the `·` is on its own line. +-/ + +/-- +The `cdot` linter flags uses of the "cdot" `·` that are achieved by typing a character +different from `·`. +For instance, a "plain" dot `.` is allowed syntax, but is flagged by the linter. +It also flags "isolated cdots", i.e. when the `·` is on its own line. -/ +register_option linter.style.cdot : Bool := { + defValue := false + descr := "enable the `cdot` linter" +} + +/-- `isCDot? stx` checks whether `stx` is a `Syntax` node corresponding to a `cdot` typed with +the character `·`. -/ +def isCDot? : Syntax → Bool + | .node _ ``cdotTk #[.node _ `patternIgnore #[.node _ _ #[.atom _ v]]] => v == "·" + | .node _ ``Lean.Parser.Term.cdot #[.atom _ v] => v == "·" + | _ => false + +/-- +`findCDot stx` extracts from `stx` the syntax nodes of `kind` `Lean.Parser.Term.cdot` or `cdotTk`. -/ +partial +def findCDot : Syntax → Array Syntax + | stx@(.node _ kind args) => + let dargs := (args.map findCDot).flatten + match kind with + | ``Lean.Parser.Term.cdot | ``cdotTk => dargs.push stx + | _ => dargs + |_ => #[] + +/-- `unwanted_cdot stx` returns an array of syntax atoms within `stx` +corresponding to `cdot`s that are not written with the character `·`. +This is precisely what the `cdot` linter flags. +-/ +def unwanted_cdot (stx : Syntax) : Array Syntax := + (findCDot stx).filter (!isCDot? ·) + +namespace Style + +@[inherit_doc linter.style.cdot] +def cdotLinter : Linter where run := withSetOptionIn fun stx ↦ do + unless Linter.getLinterValue linter.style.cdot (← getOptions) do + return + if (← MonadState.get).messages.hasErrors then + return + for s in unwanted_cdot stx do + Linter.logLint linter.style.cdot s + m!"Please, use '·' (typed as `\\.`) instead of '{s}' as 'cdot'." + -- We also check for isolated cdot's, i.e. when the cdot is on its own line. + for cdot in Mathlib.Linter.findCDot stx do + match cdot.find? (·.isOfKind `token.«· ») with + | some (.node _ _ #[.atom (.original _ _ afterCDot _) _]) => + if (afterCDot.takeWhile (·.isWhitespace)).contains '\n' then + logWarningAt cdot <| .tagged linter.style.cdot.name + m!"This central dot `·` is isolated; please merge it with the next line." + | _ => return + +initialize addLinter cdotLinter + +end Style + +/-! +# The `dollarSyntax` linter + +The `dollarSyntax` linter flags uses of `<|` that are achieved by typing `$`. +These are disallowed by the mathlib style guide, as using `<|` pairs better with `|>`. +-/ + +/-- The `dollarSyntax` linter flags uses of `<|` that are achieved by typing `$`. +These are disallowed by the mathlib style guide, as using `<|` pairs better with `|>`. -/ +register_option linter.style.dollarSyntax : Bool := { + defValue := false + descr := "enable the `dollarSyntax` linter" +} + +namespace Style.dollarSyntax + +/-- `findDollarSyntax stx` extracts from `stx` the syntax nodes of `kind` `$`. -/ +partial +def findDollarSyntax : Syntax → Array Syntax + | stx@(.node _ kind args) => + let dargs := (args.map findDollarSyntax).flatten + match kind with + | ``«term_$__» => dargs.push stx + | _ => dargs + |_ => #[] + +@[inherit_doc linter.style.dollarSyntax] +def dollarSyntaxLinter : Linter where run := withSetOptionIn fun stx ↦ do + unless Linter.getLinterValue linter.style.dollarSyntax (← getOptions) do + return + if (← MonadState.get).messages.hasErrors then + return + for s in findDollarSyntax stx do + Linter.logLint linter.style.dollarSyntax s + m!"Please use '<|' instead of '$' for the pipe operator." + +initialize addLinter dollarSyntaxLinter + +end Style.dollarSyntax + +/-! +# The `lambdaSyntax` linter + +The `lambdaSyntax` linter is a syntax linter that flags uses of the symbol `λ` to define anonymous +functions, as opposed to the `fun` keyword. These are syntactically equivalent; mathlib style +prefers the latter as it is considered more readable. +-/ + +/-- +The `lambdaSyntax` linter flags uses of the symbol `λ` to define anonymous functions. +This is syntactically equivalent to the `fun` keyword; mathlib style prefers using the latter. +-/ +register_option linter.style.lambdaSyntax : Bool := { + defValue := false + descr := "enable the `lambdaSyntax` linter" +} + +namespace Style.lambdaSyntax + +/-- +`findLambdaSyntax stx` extracts from `stx` all syntax nodes of `kind` `Term.fun`. -/ +partial +def findLambdaSyntax : Syntax → Array Syntax + | stx@(.node _ kind args) => + let dargs := (args.map findLambdaSyntax).flatten + match kind with + | ``Parser.Term.fun => dargs.push stx + | _ => dargs + |_ => #[] + +@[inherit_doc linter.style.lambdaSyntax] +def lambdaSyntaxLinter : Linter where run := withSetOptionIn fun stx ↦ do + unless Linter.getLinterValue linter.style.lambdaSyntax (← getOptions) do + return + if (← MonadState.get).messages.hasErrors then + return + for s in findLambdaSyntax stx do + if let .atom _ "λ" := s[0] then + Linter.logLint linter.style.lambdaSyntax s[0] m!"\ + Please use 'fun' and not 'λ' to define anonymous functions.\n\ + The 'λ' syntax is deprecated in mathlib4." + +initialize addLinter lambdaSyntaxLinter + +end Style.lambdaSyntax + +/-! +# The "longFile" linter + +The "longFile" linter emits a warning on files which are longer than a certain number of lines +(1500 by default). +-/ + +/-- +The "longFile" linter emits a warning on files which are longer than a certain number of lines +(`linter.style.longFileDefValue` by default on mathlib, no limit for downstream projects). +If this option is set to `N` lines, the linter warns once a file has more than `N` lines. +A value of `0` silences the linter entirely. +-/ +register_option linter.style.longFile : Nat := { + defValue := 0 + 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] +def longFileLinter : Linter where run := withSetOptionIn fun stx ↦ do + let linterBound := linter.style.longFile.get (← getOptions) + if linterBound == 0 then + return + let defValue := linter.style.longFileDefValue.get (← getOptions) + let smallOption := match stx with + | `(set_option linter.style.longFile $x) => TSyntax.getNat ⟨x.raw⟩ ≤ defValue + | _ => false + if smallOption then + logWarningAt stx <| .tagged linter.style.longFile.name + m!"The default value of the `longFile` linter is {defValue}.\n\ + The current value of {linterBound} does not exceed the allowed bound.\n\ + Please, remove the `set_option linter.style.longFile {linterBound}`." + else + -- Thanks to the above check, the linter option is either not set (and hence equal + -- to the default) or set to some value *larger* than the default. + -- `Parser.isTerminalCommand` allows `stx` to be `#exit`: this is useful for tests. + unless Parser.isTerminalCommand stx do return + -- We exclude `Mathlib.lean` from the linter: it exceeds linter's default number of allowed + -- lines, and it is an auto-generated import-only file. + -- TODO: if there are more such files, revise the implementation. + if (← getMainModule) == `Mathlib then return + if let some init := stx.getTailPos? then + -- the last line: we subtract 1, since the last line is expected to be empty + let lastLine := ((← getFileMap).toPosition init).line + -- In this case, the file has an allowed length, and the linter option is unnecessarily set. + if lastLine ≤ defValue && defValue < linterBound then + logWarningAt stx <| .tagged linter.style.longFile.name + m!"The default value of the `longFile` linter is {defValue}.\n\ + This file is {lastLine} lines long which does not exceed the allowed bound.\n\ + Please, remove the `set_option linter.style.longFile {linterBound}`." + else + -- `candidate` is divisible by `100` and satisfies `lastLine + 100 < candidate ≤ lastLine + 200` + -- note that either `lastLine ≤ defValue` and `defValue = linterBound` hold or + -- `candidate` is necessarily bigger than `lastLine` and hence bigger than `defValue` + let candidate := (lastLine / 100) * 100 + 200 + let candidate := max candidate defValue + -- In this case, the file is longer than the default and also than what the option says. + if defValue ≤ linterBound && linterBound < lastLine then + logWarningAt stx <| .tagged linter.style.longFile.name + m!"This file is {lastLine} lines long, but the limit is {linterBound}.\n\n\ + You can extend the allowed length of the file using \ + `set_option linter.style.longFile {candidate}`.\n\ + You can completely disable this linter by setting the length limit to `0`." + else + -- Finally, the file exceeds the default value, but not the option: we only allow the value + -- of the option to be `candidate` or `candidate + 100`. + -- In particular, this flags any option that is set to an unnecessarily high value. + if linterBound == candidate || linterBound + 100 == candidate then return + else + logWarningAt stx <| .tagged linter.style.longFile.name + m!"This file is {lastLine} lines long. \ + The current limit is {linterBound}, but it is expected to be {candidate}:\n\ + `set_option linter.style.longFile {candidate}`." + +initialize addLinter longFileLinter + +end Style.longFile + +/-! # The "longLine linter" -/ + +/-- The "longLine" linter emits a warning on lines longer than 100 characters. +We allow lines containing URLs to be longer, though. -/ +register_option linter.style.longLine : Bool := { + defValue := false + descr := "enable the longLine linter" +} + +namespace Style.longLine + +@[inherit_doc Mathlib.Linter.linter.style.longLine] +def longLineLinter : Linter where run := withSetOptionIn fun stx ↦ do + unless Linter.getLinterValue linter.style.longLine (← getOptions) do + return + if (← MonadState.get).messages.hasErrors then + return + -- The linter ignores the `#guard_msgs` command, in particular its doc-string. + -- The linter still lints the message guarded by `#guard_msgs`. + if stx.isOfKind ``Lean.guardMsgsCmd then + return + -- if the linter reached the end of the file, then we scan the `import` syntax instead + let stx := ← do + if stx.isOfKind ``Lean.Parser.Command.eoi then + let fileMap ← getFileMap + -- `impMods` is the syntax for the modules imported in the current file + let (impMods, _) ← Parser.parseHeader + { input := fileMap.source, fileName := ← getFileName, fileMap := fileMap } + return impMods + else return stx + let sstr := stx.getSubstring? + let fm ← getFileMap + let longLines := ((sstr.getD default).splitOn "\n").filter fun line ↦ + (100 < (fm.toPosition line.stopPos).column) + for line in longLines do + if (line.splitOn "http").length ≤ 1 then + let stringMsg := if line.contains '"' then + "\nYou can use \"string gaps\" to format long strings: within a string quotation, \ + using a '\\' at the end of a line allows you to continue the string on the following \ + line, removing all intervening whitespace." + else "" + Linter.logLint linter.style.longLine (.ofRange ⟨line.startPos, line.stopPos⟩) + m!"This line exceeds the 100 character limit, please shorten it!{stringMsg}" +initialize addLinter longLineLinter + +end Style.longLine + end Mathlib.Linter diff --git a/MathlibTest/Lint.lean b/MathlibTest/Lint.lean index 1ab5b6b71ed53..3e03fc2dae24d 100644 --- a/MathlibTest/Lint.lean +++ b/MathlibTest/Lint.lean @@ -1,7 +1,5 @@ import Mathlib.Tactic.Linter.Lint import Mathlib.Tactic.ToAdditive -import Mathlib.Order.SetNotation - /-- warning: The namespace 'add' is duplicated in the declaration 'add.add' note: this linter can be disabled with `set_option linter.dupNamespace false` @@ -56,222 +54,3 @@ note: this linter can be disabled with `set_option linter.dupNamespace false` export Nat (add add_comm add) end add - -section cdotLinter -set_option linter.style.cdot true - -set_option linter.globalAttributeIn false in -/-- -warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. -note: this linter can be disabled with `set_option linter.style.cdot false` ---- -warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. -note: this linter can be disabled with `set_option linter.style.cdot false` ---- -warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. -note: this linter can be disabled with `set_option linter.style.cdot false` --/ -#guard_msgs in -attribute [instance] Int.add in -instance : Inhabited Nat where - default := by - . have := 0 - · have : Nat → Nat → Nat := (· + .) - . exact 0 - -/-- -warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. -note: this linter can be disabled with `set_option linter.style.cdot false` --/ -#guard_msgs in -example : Add Nat where add := (. + ·) - -/-- -warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. -note: this linter can be disabled with `set_option linter.style.cdot false` --/ -#guard_msgs in -example : Add Nat where add := (. + ·) - -/-- -warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. -note: this linter can be disabled with `set_option linter.style.cdot false` ---- -warning: This central dot `·` is isolated; please merge it with the next line. ---- -warning: This central dot `·` is isolated; please merge it with the next line. --/ -#guard_msgs in -example : Nat := by - have : Nat := by - · - -- some empty have - have := 0 - · - - -- another - have := 1 - . exact 2 - exact 0 - -#guard_msgs in -example : True := by - have : Nat := by - -- This is how code should look: no error. - · -- comment - exact 37 - trivial - -end cdotLinter -set_option linter.style.dollarSyntax true - -set_option linter.globalAttributeIn false in -/-- -warning: Please use '<|' instead of '$' for the pipe operator. -note: this linter can be disabled with `set_option linter.style.dollarSyntax false` ---- -warning: Please use '<|' instead of '$' for the pipe operator. -note: this linter can be disabled with `set_option linter.style.dollarSyntax false` --/ -#guard_msgs in -attribute [instance] Int.add in -instance (f g : Nat → Nat) : Inhabited Nat where - default := by - · have := 0 - · have : Nat := f $ g $ 0 - · exact 0 - -section lambdaSyntaxLinter - -set_option linter.style.lambdaSyntax true - -/-- -warning: Please use 'fun' and not 'λ' to define anonymous functions. -The 'λ' syntax is deprecated in mathlib4. -note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` --/ -#guard_msgs in -example : ℕ → ℕ := λ _ ↦ 0 - -/-- -warning: Please use 'fun' and not 'λ' to define anonymous functions. -The 'λ' syntax is deprecated in mathlib4. -note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` --/ -#guard_msgs in -def foo : Bool := by - let _f : ℕ → ℕ := λ _ ↦ 0 - exact true - -example : ℕ → ℕ := fun n ↦ n - 1 - -/-- -warning: Please use 'fun' and not 'λ' to define anonymous functions. -The 'λ' syntax is deprecated in mathlib4. -note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` --/ -#guard_msgs in -example : ℕ → ℕ := by exact λ n ↦ 3 * n + 1 - -/-- -warning: declaration uses 'sorry' ---- -warning: Please use 'fun' and not 'λ' to define anonymous functions. -The 'λ' syntax is deprecated in mathlib4. -note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` ---- -warning: Please use 'fun' and not 'λ' to define anonymous functions. -The 'λ' syntax is deprecated in mathlib4. -note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` ---- -warning: Please use 'fun' and not 'λ' to define anonymous functions. -The 'λ' syntax is deprecated in mathlib4. -note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` --/ -#guard_msgs in -example : ℕ → ℕ → ℕ → ℕ := by - have (n : ℕ) : True := trivial - have : (Set.univ : Set ℕ) = ⋃ (i : ℕ), (Set.iUnion λ j ↦ ({0, j} : Set ℕ)) := sorry - have : ∃ m : ℕ, ⋃ i : ℕ, (Set.univ : Set ℕ) = ∅ := sorry - exact λ _a ↦ fun _b ↦ λ _c ↦ 0 - -/-- -warning: Please use 'fun' and not 'λ' to define anonymous functions. -The 'λ' syntax is deprecated in mathlib4. -note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` ---- -warning: Please use 'fun' and not 'λ' to define anonymous functions. -The 'λ' syntax is deprecated in mathlib4. -note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` ---- -warning: Please use 'fun' and not 'λ' to define anonymous functions. -The 'λ' syntax is deprecated in mathlib4. -note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` --/ -#guard_msgs in -example : True := by - have : 0 = 0 ∧ 0 = 0 ∧ 1 + 3 = 4 := by - refine ⟨by trivial, by - let _f := λ n : ℕ ↦ 0; - have : ℕ := by - · -- comment - · have := λ k : ℕ ↦ -5 - · exact 0 - refine ⟨by trivial, have := λ k : ℕ ↦ -5; by simp⟩ - ⟩ - trivial - --- Code such as the following would require walking the infotree instead: --- the inner set_option is ignore (in either direction). --- As this seems unlikely to occur by accident and its use is dubious, we don't worry about this. -/-- -warning: Please use 'fun' and not 'λ' to define anonymous functions. -The 'λ' syntax is deprecated in mathlib4. -note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` --/ -#guard_msgs in -example : ℕ → ℕ := λ _ ↦ 0 - -set_option linter.style.lambdaSyntax false -#guard_msgs in -example : ℕ → ℕ := λ _ ↦ 0 - -end lambdaSyntaxLinter - -set_option linter.style.longLine true -/-- -warning: This line exceeds the 100 character limit, please shorten it! -note: this linter can be disabled with `set_option linter.style.longLine false` --/ -#guard_msgs in -/-! -/ - -#guard_msgs in --- Lines with more than 100 characters containing URLs are allowed. -/-! http -/ - -set_option linter.style.longLine true --- The *argument* of `#guard_msgs` is *not* exempt from the linter. -/-- -warning: This line exceeds the 100 character limit, please shorten it! -note: this linter can be disabled with `set_option linter.style.longLine false` --/ -#guard_msgs in #guard true - --- However, the *doc-string* of #guard_msgs is exempt from the linter: --- these are automatically generated, hence linting them is not helpful. -/-- -info: [0, 1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21, 22, 23, 24, 25, 26] --/ -#guard_msgs in -#eval List.range 27 - -/-- -info: " \" " : String ---- -warning: This line exceeds the 100 character limit, please shorten it! -You can use "string gaps" to format long strings: within a string quotation, using a '\' at the end of a line allows you to continue the string on the following line, removing all intervening whitespace. -note: this linter can be disabled with `set_option linter.style.longLine false` --/ -#guard_msgs in -#check " \" " diff --git a/MathlibTest/LintStyle.lean b/MathlibTest/LintStyle.lean index 82826979b6f8c..719fd833552a4 100644 --- a/MathlibTest/LintStyle.lean +++ b/MathlibTest/LintStyle.lean @@ -1,5 +1,5 @@ import Mathlib.Tactic.Linter.Style -import Mathlib.Tactic.Common +import Mathlib.Order.SetNotation /-! Tests for all the style linters. -/ @@ -120,3 +120,219 @@ lemma foo' : True := trivial -- TODO: add terms for the term form end setOption + +section cdotLinter +set_option linter.style.cdot true + +set_option linter.globalAttributeIn false in +/-- +warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. +note: this linter can be disabled with `set_option linter.style.cdot false` +--- +warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. +note: this linter can be disabled with `set_option linter.style.cdot false` +--- +warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. +note: this linter can be disabled with `set_option linter.style.cdot false` +-/ +#guard_msgs in +attribute [instance] Int.add in +instance : Inhabited Nat where + default := by + . have := 0 + · have : Nat → Nat → Nat := (· + .) + . exact 0 + +/-- +warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. +note: this linter can be disabled with `set_option linter.style.cdot false` +-/ +#guard_msgs in +example : Add Nat where add := (. + ·) + +/-- +warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. +note: this linter can be disabled with `set_option linter.style.cdot false` +-/ +#guard_msgs in +example : Add Nat where add := (. + ·) + +/-- +warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. +note: this linter can be disabled with `set_option linter.style.cdot false` +--- +warning: This central dot `·` is isolated; please merge it with the next line. +--- +warning: This central dot `·` is isolated; please merge it with the next line. +-/ +#guard_msgs in +example : Nat := by + have : Nat := by + · + -- some empty have + have := 0 + · + + -- another + have := 1 + . exact 2 + exact 0 + +#guard_msgs in +example : True := by + have : Nat := by + -- This is how code should look: no error. + · -- comment + exact 37 + trivial + +end cdotLinter +set_option linter.style.dollarSyntax true + +set_option linter.globalAttributeIn false in +/-- +warning: Please use '<|' instead of '$' for the pipe operator. +note: this linter can be disabled with `set_option linter.style.dollarSyntax false` +--- +warning: Please use '<|' instead of '$' for the pipe operator. +note: this linter can be disabled with `set_option linter.style.dollarSyntax false` +-/ +#guard_msgs in +attribute [instance] Int.add in +instance (f g : Nat → Nat) : Inhabited Nat where + default := by + · have := 0 + · have : Nat := f $ g $ 0 + · exact 0 + +section lambdaSyntaxLinter + +set_option linter.style.lambdaSyntax true + +/-- +warning: Please use 'fun' and not 'λ' to define anonymous functions. +The 'λ' syntax is deprecated in mathlib4. +note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` +-/ +#guard_msgs in +example : ℕ → ℕ := λ _ ↦ 0 + +/-- +warning: Please use 'fun' and not 'λ' to define anonymous functions. +The 'λ' syntax is deprecated in mathlib4. +note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` +-/ +#guard_msgs in +def foo : Bool := by + let _f : ℕ → ℕ := λ _ ↦ 0 + exact true + +example : ℕ → ℕ := fun n ↦ n - 1 + +/-- +warning: Please use 'fun' and not 'λ' to define anonymous functions. +The 'λ' syntax is deprecated in mathlib4. +note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` +-/ +#guard_msgs in +example : ℕ → ℕ := by exact λ n ↦ 3 * n + 1 + +/-- +warning: declaration uses 'sorry' +--- +warning: Please use 'fun' and not 'λ' to define anonymous functions. +The 'λ' syntax is deprecated in mathlib4. +note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` +--- +warning: Please use 'fun' and not 'λ' to define anonymous functions. +The 'λ' syntax is deprecated in mathlib4. +note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` +--- +warning: Please use 'fun' and not 'λ' to define anonymous functions. +The 'λ' syntax is deprecated in mathlib4. +note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` +-/ +#guard_msgs in +example : ℕ → ℕ → ℕ → ℕ := by + have (n : ℕ) : True := trivial + have : (Set.univ : Set ℕ) = ⋃ (i : ℕ), (Set.iUnion λ j ↦ ({0, j} : Set ℕ)) := sorry + have : ∃ m : ℕ, ⋃ i : ℕ, (Set.univ : Set ℕ) = ∅ := sorry + exact λ _a ↦ fun _b ↦ λ _c ↦ 0 + +/-- +warning: Please use 'fun' and not 'λ' to define anonymous functions. +The 'λ' syntax is deprecated in mathlib4. +note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` +--- +warning: Please use 'fun' and not 'λ' to define anonymous functions. +The 'λ' syntax is deprecated in mathlib4. +note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` +--- +warning: Please use 'fun' and not 'λ' to define anonymous functions. +The 'λ' syntax is deprecated in mathlib4. +note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` +-/ +#guard_msgs in +example : True := by + have : 0 = 0 ∧ 0 = 0 ∧ 1 + 3 = 4 := by + refine ⟨by trivial, by + let _f := λ n : ℕ ↦ 0; + have : ℕ := by + · -- comment + · have := λ k : ℕ ↦ -5 + · exact 0 + refine ⟨by trivial, have := λ k : ℕ ↦ -5; by simp⟩ + ⟩ + trivial + +-- Code such as the following would require walking the infotree instead: +-- the inner set_option is ignore (in either direction). +-- As this seems unlikely to occur by accident and its use is dubious, we don't worry about this. +/-- +warning: Please use 'fun' and not 'λ' to define anonymous functions. +The 'λ' syntax is deprecated in mathlib4. +note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` +-/ +#guard_msgs in +example : ℕ → ℕ := λ _ ↦ 0 + +set_option linter.style.lambdaSyntax false +#guard_msgs in +example : ℕ → ℕ := λ _ ↦ 0 + +end lambdaSyntaxLinter + +set_option linter.style.longLine true +/-- +warning: This line exceeds the 100 character limit, please shorten it! +note: this linter can be disabled with `set_option linter.style.longLine false` +-/ +#guard_msgs in +/-! -/ + +#guard_msgs in +-- Lines with more than 100 characters containing URLs are allowed. +/-! http -/ + +set_option linter.style.longLine true +-- The *argument* of `#guard_msgs` is *not* exempt from the linter. +/-- +warning: This line exceeds the 100 character limit, please shorten it! +note: this linter can be disabled with `set_option linter.style.longLine false` +-/ +#guard_msgs in #guard true + +-- However, the *doc-string* of #guard_msgs is exempt from the linter: +-- these are automatically generated, hence linting them is not helpful. +/-- +info: [0, 1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21, 22, 23, 24, 25, 26] +-/ +#guard_msgs in +#eval List.range 27 + +-- TODO: this used to error about the 100 character limit (mentioning string gaps), +-- restore this test! +/-- info: " \" " : String -/ +#guard_msgs in +#check " \" \ + " diff --git a/MathlibTest/LongFile.lean b/MathlibTest/LongFile.lean index 635f2ab485c30..35b772bc6c409 100644 --- a/MathlibTest/LongFile.lean +++ b/MathlibTest/LongFile.lean @@ -1,4 +1,4 @@ -import Mathlib.Tactic.Linter.Lint +import Mathlib.Tactic.Linter.Style /- # Testing the `longFile` linter From 8489bbe7ee4cb7bea53236fda9a68368f16d5364 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Wed, 27 Nov 2024 12:53:27 +0000 Subject: [PATCH 354/829] chore: update Mathlib dependencies 2024-11-27 (#19536) 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 a3160f181a6cb..57203830de6b6 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "f6d16c2a073e0385a362ad3e5d9e7e8260e298d6", + "rev": "c933dd9b00271d869e22b802a015092d1e8e454a", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From 37c90e46ca44c28b3c963d680ad3f4107279cbc1 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Wed, 27 Nov 2024 13:34:02 +0000 Subject: [PATCH 355/829] feat: uniform time lemma for the existence of global integral curves (#9013) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Lemma 9.15 in Lee's Introduction to Smooth Manifolds: > Let `v` be a smooth vector field on a smooth manifold `M`. If there exists `ε > 0` such that for each point `x : M`, there exists an integral curve of `v` through `x` defined on an open interval `Ioo (-ε) ε`, then every point on `M` has a global integral curve of `v` passing through it. We only require `v` to be $C^1$. To achieve this, we define the extension of an integral curve `γ` by another integral curve `γ'`, if they agree at a point inside their overlapping open interval domains. This utilises the uniqueness theorem of integral curves. We need this lemma to show that vector fields on compact manifolds always have global integral curves. - [x] depends on: #8886 - [x] depends on: #18833 - [x] depends on: #19008 Co-authored-by: Michael Rothgang Co-authored-by: Johan Commelin --- Mathlib.lean | 1 + .../Manifold/IntegralCurve/ExistUnique.lean | 4 +- .../Manifold/IntegralCurve/UniformTime.lean | 210 ++++++++++++++++++ Mathlib/Order/Interval/Set/Basic.lean | 8 +- 4 files changed, 218 insertions(+), 5 deletions(-) create mode 100644 Mathlib/Geometry/Manifold/IntegralCurve/UniformTime.lean diff --git a/Mathlib.lean b/Mathlib.lean index a843c7987f61d..f3191d04fb72c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3047,6 +3047,7 @@ import Mathlib.Geometry.Manifold.Instances.UnitsOfNormedAlgebra import Mathlib.Geometry.Manifold.IntegralCurve.Basic import Mathlib.Geometry.Manifold.IntegralCurve.ExistUnique import Mathlib.Geometry.Manifold.IntegralCurve.Transform +import Mathlib.Geometry.Manifold.IntegralCurve.UniformTime import Mathlib.Geometry.Manifold.InteriorBoundary import Mathlib.Geometry.Manifold.LocalDiffeomorph import Mathlib.Geometry.Manifold.LocalInvariantProperties diff --git a/Mathlib/Geometry/Manifold/IntegralCurve/ExistUnique.lean b/Mathlib/Geometry/Manifold/IntegralCurve/ExistUnique.lean index 369e3803fc7bd..8960304376097 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve/ExistUnique.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve/ExistUnique.lean @@ -250,7 +250,7 @@ theorem isIntegralCurve_Ioo_eq_of_contMDiff_boundaryless [BoundarylessManifold I period `a - b`. -/ lemma IsIntegralCurve.periodic_of_eq [BoundarylessManifold I M] (hγ : IsIntegralCurve γ v) - (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) + (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) (heq : γ a = γ b) : Periodic γ (a - b) := by intro t apply congrFun <| @@ -260,7 +260,7 @@ lemma IsIntegralCurve.periodic_of_eq [BoundarylessManifold I M] /-- A global integral curve is injective xor periodic with positive period. -/ lemma IsIntegralCurve.periodic_xor_injective [BoundarylessManifold I M] (hγ : IsIntegralCurve γ v) - (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) : + (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) : Xor' (∃ T > 0, Periodic γ T) (Injective γ) := by rw [xor_iff_iff_not] refine ⟨fun ⟨T, hT, hf⟩ ↦ hf.not_injective (ne_of_gt hT), ?_⟩ diff --git a/Mathlib/Geometry/Manifold/IntegralCurve/UniformTime.lean b/Mathlib/Geometry/Manifold/IntegralCurve/UniformTime.lean new file mode 100644 index 0000000000000..6d11987df7df6 --- /dev/null +++ b/Mathlib/Geometry/Manifold/IntegralCurve/UniformTime.lean @@ -0,0 +1,210 @@ +/- +Copyright (c) 2023 Winston Yin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Winston Yin +-/ +import Mathlib.Geometry.Manifold.IntegralCurve.ExistUnique + +/-! +# Uniform time lemma for the global existence of integral curves + +## Main results + +* `exists_isIntegralCurve_of_isIntegralCurveOn`: If there exists `ε > 0` such that the local +integral curve at each point `x : M` is defined at least on an open interval `Ioo (-ε) ε`, then +every point on `M` has a global integral curve passing through it. + +## Reference + +* [Lee, J. M. (2012). _Introduction to Smooth Manifolds_. Springer New York.][lee2012] + +## Tags + +integral curve, vector field, global existence +-/ + +open scoped Topology + +open Function Set + +variable + {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] + {H : Type*} [TopologicalSpace H] {I : ModelWithCorners ℝ E H} + {M : Type*} [TopologicalSpace M] [ChartedSpace H M] [SmoothManifoldWithCorners I M] [T2Space M] + {γ γ' : ℝ → M} {v : (x : M) → TangentSpace I x} {s s' : Set ℝ} {t₀ : ℝ} + +/-- This is the uniqueness theorem of integral curves applied to a real-indexed family of integral + curves with the same starting point. -/ +lemma eqOn_of_isIntegralCurveOn_Ioo [BoundarylessManifold I M] + (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) {x : M} + (γ : ℝ → ℝ → M) (hγx : ∀ a, γ a 0 = x) (hγ : ∀ a > 0, IsIntegralCurveOn (γ a) v (Ioo (-a) a)) + {a a' : ℝ} (hpos : 0 < a') (hle : a' ≤ a) : + EqOn (γ a') (γ a) (Ioo (-a') a') := by + apply isIntegralCurveOn_Ioo_eqOn_of_contMDiff_boundaryless _ hv + (hγ a' (by positivity)) ((hγ a (gt_of_ge_of_gt hle hpos)).mono _) + (by rw [hγx a, hγx a']) + · rw [mem_Ioo] + exact ⟨neg_lt_zero.mpr hpos, by positivity⟩ + · apply Ioo_subset_Ioo <;> linarith + +/-- For a family of integral curves `γ : ℝ → ℝ → M` with the same starting point `γ 0 = x` such that + each `γ a` is defined on `Ioo (-a) a`, the global curve `γ_ext := fun t ↦ γ (|t| + 1) t` agrees + with each `γ a` on `Ioo (-a) a`. This will help us show that `γ_ext` is a global integral + curve. -/ +lemma eqOn_abs_add_one_of_isIntegralCurveOn_Ioo [BoundarylessManifold I M] + (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) {x : M} + (γ : ℝ → ℝ → M) (hγx : ∀ a, γ a 0 = x) (hγ : ∀ a > 0, IsIntegralCurveOn (γ a) v (Ioo (-a) a)) + {a : ℝ} : EqOn (fun t ↦ γ (|t| + 1) t) (γ a) (Ioo (-a) a) := by + intros t ht + by_cases hlt : |t| + 1 < a + · exact eqOn_of_isIntegralCurveOn_Ioo hv γ hγx hγ + (by positivity) hlt.le (abs_lt.mp <| lt_add_one _) + · exact eqOn_of_isIntegralCurveOn_Ioo hv γ hγx hγ + (neg_lt_self_iff.mp <| lt_trans ht.1 ht.2) (not_lt.mp hlt) ht |>.symm + +/-- For a family of integral curves `γ : ℝ → ℝ → M` with the same starting point `γ 0 = x` such that + each `γ a` is defined on `Ioo (-a) a`, the function `γ_ext := fun t ↦ γ (|t| + 1) t` is a global + integral curve. -/ +lemma isIntegralCurve_abs_add_one_of_isIntegralCurveOn_Ioo [BoundarylessManifold I M] + (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) {x : M} + (γ : ℝ → ℝ → M) (hγx : ∀ a, γ a 0 = x) (hγ : ∀ a > 0, IsIntegralCurveOn (γ a) v (Ioo (-a) a)) : + IsIntegralCurve (fun t ↦ γ (|t| + 1) t) v := by + intro t + apply HasMFDerivAt.congr_of_eventuallyEq (f := γ (|t| + 1)) + · apply hγ (|t| + 1) (by positivity) + rw [mem_Ioo, ← abs_lt] + exact lt_add_one _ + · rw [Filter.eventuallyEq_iff_exists_mem] + refine ⟨Ioo (-(|t| + 1)) (|t| + 1), ?_, + eqOn_abs_add_one_of_isIntegralCurveOn_Ioo hv γ hγx hγ⟩ + have : |t| < |t| + 1 := lt_add_of_pos_right |t| zero_lt_one + rw [abs_lt] at this + exact Ioo_mem_nhds this.1 this.2 + +/-- The existence of a global integral curve is equivalent to the existence of a family of local + integral curves `γ : ℝ → ℝ → M` with the same starting point `γ 0 = x` such that each `γ a` is + defined on `Ioo (-a) a`. -/ +lemma exists_isIntegralCurve_iff_exists_isIntegralCurveOn_Ioo [BoundarylessManifold I M] + (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) (x : M) : + (∃ γ, γ 0 = x ∧ IsIntegralCurve γ v) ↔ + ∀ a, ∃ γ, γ 0 = x ∧ IsIntegralCurveOn γ v (Ioo (-a) a) := by + refine ⟨fun ⟨γ, h1, h2⟩ _ ↦ ⟨γ, h1, h2.isIntegralCurveOn _⟩, fun h ↦ ?_⟩ + choose γ hγx hγ using h + exact ⟨fun t ↦ γ (|t| + 1) t, hγx (|0| + 1), + isIntegralCurve_abs_add_one_of_isIntegralCurveOn_Ioo hv γ hγx (fun a _ ↦ hγ a)⟩ + +/-- Let `γ` and `γ'` be integral curves defined on `Ioo a b` and `Ioo a' b'`, respectively. Then, + `piecewise (Ioo a b) γ γ'` is equal to `γ` and `γ'` in their respective domains. + `Set.piecewise_eqOn` shows the equality for `γ` by definition, while this lemma shows the equality + for `γ'` by the uniqueness of integral curves. -/ +lemma eqOn_piecewise_of_isIntegralCurveOn_Ioo [BoundarylessManifold I M] + (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) + {a b a' b' : ℝ} (hγ : IsIntegralCurveOn γ v (Ioo a b)) + (hγ' : IsIntegralCurveOn γ' v (Ioo a' b')) + (ht₀ : t₀ ∈ Ioo a b ∩ Ioo a' b') (h : γ t₀ = γ' t₀) : + EqOn (piecewise (Ioo a b) γ γ') γ' (Ioo a' b') := by + intros t ht + suffices H : EqOn γ γ' (Ioo (max a a') (min b b')) by + by_cases hmem : t ∈ Ioo a b + · rw [piecewise, if_pos hmem] + apply H + simp [ht.1, ht.2, hmem.1, hmem.2] + · rw [piecewise, if_neg hmem] + apply isIntegralCurveOn_Ioo_eqOn_of_contMDiff_boundaryless _ hv + (hγ.mono (Ioo_subset_Ioo (le_max_left ..) (min_le_left ..))) + (hγ'.mono (Ioo_subset_Ioo (le_max_right ..) (min_le_right ..))) h + exact ⟨max_lt ht₀.1.1 ht₀.2.1, lt_min ht₀.1.2 ht₀.2.2⟩ + +/-- The extension of an integral curve by another integral curve is an integral curve. + + If two integral curves are defined on overlapping open intervals, and they agree at a point in + their common domain, then they can be patched together to form a longer integral curve. + + This is stated for manifolds without boundary for simplicity. We actually only need to assume that + the images of `γ` and `γ'` lie in the interior of the manifold. TODO: Generalise to manifolds with + boundary. -/ +lemma isIntegralCurveOn_piecewise [BoundarylessManifold I M] + (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) + {a b a' b' : ℝ} (hγ : IsIntegralCurveOn γ v (Ioo a b)) + (hγ' : IsIntegralCurveOn γ' v (Ioo a' b')) {t₀ : ℝ} + (ht₀ : t₀ ∈ Ioo a b ∩ Ioo a' b') (h : γ t₀ = γ' t₀) : + IsIntegralCurveOn (piecewise (Ioo a b) γ γ') v (Ioo a b ∪ Ioo a' b') := by + intros t ht + by_cases hmem : t ∈ Ioo a b + · rw [piecewise, if_pos hmem] + apply (hγ t hmem).congr_of_eventuallyEq + rw [Filter.eventuallyEq_iff_exists_mem] + refine ⟨Ioo a b, isOpen_Ioo.mem_nhds hmem, ?_⟩ + intros t' ht' + rw [piecewise, if_pos ht'] + · rw [mem_union, or_iff_not_imp_left] at ht + rw [piecewise, if_neg hmem] + apply (hγ' t <| ht hmem).congr_of_eventuallyEq + rw [Filter.eventuallyEq_iff_exists_mem] + exact ⟨Ioo a' b', isOpen_Ioo.mem_nhds <| ht hmem, + eqOn_piecewise_of_isIntegralCurveOn_Ioo hv hγ hγ' ht₀ h⟩ + +/-- If there exists `ε > 0` such that the local integral curve at each point `x : M` is defined at + least on an open interval `Ioo (-ε) ε`, then every point on `M` has a global integral + curve passing through it. + + See Lemma 9.15, [J.M. Lee (2012)][lee2012]. -/ +lemma exists_isIntegralCurve_of_isIntegralCurveOn [BoundarylessManifold I M] + {v : (x : M) → TangentSpace I x} + (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) + {ε : ℝ} (hε : 0 < ε) (h : ∀ x : M, ∃ γ : ℝ → M, γ 0 = x ∧ IsIntegralCurveOn γ v (Ioo (-ε) ε)) + (x : M) : ∃ γ : ℝ → M, γ 0 = x ∧ IsIntegralCurve γ v := by + let s := { a | ∃ γ, γ 0 = x ∧ IsIntegralCurveOn γ v (Ioo (-a) a) } + suffices hbdd : ¬BddAbove s by + rw [not_bddAbove_iff] at hbdd + rw [exists_isIntegralCurve_iff_exists_isIntegralCurveOn_Ioo hv] + intro a + obtain ⟨y, ⟨γ, hγ1, hγ2⟩, hlt⟩ := hbdd a + exact ⟨γ, hγ1, hγ2.mono <| Ioo_subset_Ioo (neg_le_neg hlt.le) hlt.le⟩ + intro hbdd + set asup := sSup s with hasup + -- we will obtain two integral curves, one centred at some `t₀ > 0` with + -- `0 ≤ asup - ε < t₀ < asup`; let `t₀ = asup - ε / 2` + -- another centred at 0 with domain up to `a ∈ S` with `t₀ < a < asup` + obtain ⟨a, ha, hlt⟩ := Real.add_neg_lt_sSup (⟨ε, h x⟩ : Set.Nonempty s) (ε := - (ε / 2)) + (by rw [neg_lt, neg_zero]; exact half_pos hε) + rw [mem_setOf] at ha + rw [← hasup, ← sub_eq_add_neg] at hlt + + -- integral curve defined on `Ioo (-a) a` + obtain ⟨γ, h0, hγ⟩ := ha + -- integral curve starting at `-(asup - ε / 2)` with radius `ε` + obtain ⟨γ1_aux, h1_aux, hγ1⟩ := h (γ (-(asup - ε / 2))) + rw [isIntegralCurveOn_comp_add (dt := asup - ε / 2)] at hγ1 + set γ1 := γ1_aux ∘ (· + (asup - ε / 2)) with γ1_def + have heq1 : γ1 (-(asup - ε / 2)) = γ (-(asup - ε / 2)) := by simp [γ1_def, h1_aux] + -- integral curve starting at `asup - ε / 2` with radius `ε` + obtain ⟨γ2_aux, h2_aux, hγ2⟩ := h (γ (asup - ε / 2)) + rw [isIntegralCurveOn_comp_sub (dt := asup - ε / 2)] at hγ2 + set γ2 := γ2_aux ∘ (· - (asup - ε / 2)) with γ2_def + have heq2 : γ2 (asup - ε / 2) = γ (asup - ε / 2) := by simp [γ2_def, h2_aux] + + -- rewrite shifted Ioo as Ioo + rw [neg_sub] at hγ1 + rw [Real.Ioo_eq_ball, neg_add_cancel, zero_div, sub_neg_eq_add, add_self_div_two, + Metric.vadd_ball, vadd_eq_add, add_zero, Real.ball_eq_Ioo] at hγ1 hγ2 + + -- to help `linarith` + have hεle : ε ≤ asup := le_csSup hbdd (h x) + + -- extend `γ` on the left by `γ1` and on the right by `γ2` + set γ_ext : ℝ → M := piecewise (Ioo (-(asup + ε / 2)) a) + (piecewise (Ioo (-a) a) γ γ1) γ2 with γ_ext_def + have heq_ext : γ_ext 0 = x := by + rw [γ_ext_def, piecewise, if_pos ⟨by linarith, by linarith⟩, piecewise, + if_pos ⟨by linarith, by linarith⟩, h0] + -- `asup + ε / 2` is an element of `s` greater than `asup`, a contradiction + suffices hext : IsIntegralCurveOn γ_ext v (Ioo (-(asup + ε / 2)) (asup + ε / 2)) from + (not_lt.mpr <| le_csSup hbdd ⟨γ_ext, heq_ext, hext⟩) <| lt_add_of_pos_right asup (half_pos hε) + apply (isIntegralCurveOn_piecewise (t₀ := asup - ε / 2) hv _ hγ2 + ⟨⟨by linarith, hlt⟩, ⟨by linarith, by linarith⟩⟩ + (by rw [piecewise, if_pos ⟨by linarith, hlt⟩, ← heq2])).mono + (Ioo_subset_Ioo_union_Ioo le_rfl (by linarith) (by linarith)) + exact (isIntegralCurveOn_piecewise (t₀ := -(asup - ε / 2)) hv hγ hγ1 + ⟨⟨neg_lt_neg hlt, by linarith⟩, ⟨by linarith, by linarith⟩⟩ heq1.symm).mono + (union_comm _ _ ▸ Ioo_subset_Ioo_union_Ioo (by linarith) (by linarith) le_rfl) diff --git a/Mathlib/Order/Interval/Set/Basic.lean b/Mathlib/Order/Interval/Set/Basic.lean index 7b2b1871c84e9..fb3aefe5d8181 100644 --- a/Mathlib/Order/Interval/Set/Basic.lean +++ b/Mathlib/Order/Interval/Set/Basic.lean @@ -1125,7 +1125,6 @@ theorem Iic_union_Ico_eq_Iio (h : a < b) : Iic a ∪ Ico a b = Iio b := /-! #### Two finite intervals, `I?o` and `Ic?` -/ - theorem Ioo_subset_Ioo_union_Ico : Ioo a c ⊆ Ioo a b ∪ Ico b c := fun x hx => (lt_or_le x b).elim (fun hxb => Or.inl ⟨hx.1, hxb⟩) fun hxb => Or.inr ⟨hxb, hx.2⟩ @@ -1181,7 +1180,6 @@ theorem Ioo_union_Icc_eq_Ioc (h₁ : a < b) (h₂ : b ≤ c) : Ioo a b ∪ Icc b /-! #### Two finite intervals, `I?c` and `Io?` -/ - theorem Ioo_subset_Ioc_union_Ioo : Ioo a c ⊆ Ioc a b ∪ Ioo b c := fun x hx => (le_or_lt x b).elim (fun hxb => Or.inl ⟨hx.1, hxb⟩) fun hxb => Or.inr ⟨hxb, hx.2⟩ @@ -1237,7 +1235,6 @@ theorem Ioc_union_Ioc (h₁ : min a b ≤ max c d) (h₂ : min c d ≤ max a b) /-! #### Two finite intervals with a common point -/ - theorem Ioo_subset_Ioc_union_Ico : Ioo a c ⊆ Ioc a b ∪ Ico b c := Subset.trans Ioo_subset_Ioc_union_Ioo (union_subset_union_right _ Ioo_subset_Ico_self) @@ -1317,6 +1314,11 @@ theorem Ioo_union_Ioo (h₁ : min a b < max c d) (h₂ : min c d < max a b) : simp [*, min_eq_left_of_lt, min_eq_right_of_lt, max_eq_left_of_lt, max_eq_right_of_lt, le_of_lt h₂, le_of_lt h₁] +theorem Ioo_subset_Ioo_union_Ioo (h₁ : a ≤ a₁) (h₂ : c < b) (h₃ : b₁ ≤ d) : + Ioo a₁ b₁ ⊆ Ioo a b ∪ Ioo c d := fun x hx => + (lt_or_le x b).elim (fun hxb => Or.inl ⟨lt_of_le_of_lt h₁ hx.1, hxb⟩) + fun hxb => Or.inr ⟨lt_of_lt_of_le h₂ hxb, lt_of_lt_of_le hx.2 h₃⟩ + end LinearOrder section Lattice From af58035408a3fc62e9ad7e709a01d4f5f97cd714 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 27 Nov 2024 14:12:32 +0000 Subject: [PATCH 356/829] chore(discover-lean-pr-testing): add PR number and diff link to relevant branches (#19532) --- .github/workflows/discover-lean-pr-testing.yml | 9 +++++++-- 1 file changed, 7 insertions(+), 2 deletions(-) diff --git a/.github/workflows/discover-lean-pr-testing.yml b/.github/workflows/discover-lean-pr-testing.yml index 349945a0ee8fe..fa3dcad9ae5fe 100644 --- a/.github/workflows/discover-lean-pr-testing.yml +++ b/.github/workflows/discover-lean-pr-testing.yml @@ -100,8 +100,13 @@ jobs: # Check if the diff contains files other than the specified ones if echo "$DIFF_FILES" | grep -v -e 'lake-manifest.json' -e 'lakefile.lean' -e 'lean-toolchain'; then - # If it does, add the branch to RELEVANT_BRANCHES - RELEVANT_BRANCHES="$RELEVANT_BRANCHES"$'\n- '"$BRANCH" + # Extract the PR number and actual branch name + PR_NUMBER=$(echo "$BRANCH" | grep -oP '\d+$') + ACTUAL_BRANCH=${BRANCH#origin/} + GITHUB_DIFF="https://github.com/leanprover-community/mathlib4/compare/nightly-testing...$ACTUAL_BRANCH" + + # Append the branch details to RELEVANT_BRANCHES + RELEVANT_BRANCHES="$RELEVANT_BRANCHES"$'\n- '"$ACTUAL_BRANCH (lean4#$PR_NUMBER) ([diff with \`nightly-testing\`]($GITHUB_DIFF))" fi done From 8390aa74f64d70fa884b2591c9d3bf9e54ebb3fe Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 27 Nov 2024 14:12:34 +0000 Subject: [PATCH 357/829] chore(zulip_emoji_merge_delegate): exclude rss channel in python logic instead of zulip search (#19540) --- scripts/zulip_emoji_merge_delegate.py | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/scripts/zulip_emoji_merge_delegate.py b/scripts/zulip_emoji_merge_delegate.py index 91836a41e49ff..f31e66c7b506f 100644 --- a/scripts/zulip_emoji_merge_delegate.py +++ b/scripts/zulip_emoji_merge_delegate.py @@ -23,14 +23,14 @@ site=ZULIP_SITE ) -# Fetch the messages containing the PR number from the public channels (exluding #rss). +# Fetch the messages containing the PR number from the public channels. # There does not seem to be a way to search simultaneously public and private channels. public_response = client.get_messages({ "anchor": "newest", "num_before": 5000, "num_after": 0, "narrow": [ - {"operator": "channel", "operand": "rss", "negated": True}, + {"operator": "channels", "operand": "public"}, {"operator": "search", "operand": f'#{PR_NUMBER}'}, ], }) @@ -57,6 +57,8 @@ print(f"Searching for: '{pr_pattern}'") for message in messages: + if message['display_recipient'] == 'rss': + continue content = message['content'] # Check for emoji reactions reactions = message['reactions'] From b5fb043af905e8b3bb5837dafc9b0b55d12e7092 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 27 Nov 2024 15:22:44 +0000 Subject: [PATCH 358/829] feat(Fin/Tuple): add `Fin.update_insertNth` (#19542) --- Mathlib/Data/Fin/Tuple/Basic.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/Mathlib/Data/Fin/Tuple/Basic.lean b/Mathlib/Data/Fin/Tuple/Basic.lean index dd1645bcb7380..430153e9bdb75 100644 --- a/Mathlib/Data/Fin/Tuple/Basic.lean +++ b/Mathlib/Data/Fin/Tuple/Basic.lean @@ -949,6 +949,12 @@ open Set lemma insertNth_self_removeNth (p : Fin (n + 1)) (f : ∀ j, α j) : insertNth p (f p) (removeNth p f) = f := by simp +@[simp] +theorem update_insertNth (p : Fin (n + 1)) (x y : α p) (f : ∀ i, α (p.succAbove i)) : + update (p.insertNth x f) p y = p.insertNth y f := by + ext i + cases i using p.succAboveCases <;> simp [succAbove_ne] + /-- Equivalence between tuples of length `n + 1` and pairs of an element and a tuple of length `n` given by separating out the `p`-th element of the tuple. From 6a0cdd15b083f05786bf1de0eb7df4c57d925d6e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Wed, 27 Nov 2024 18:07:36 +0000 Subject: [PATCH 359/829] feat(Algebra/Module): a presentation of an algebra induces a presentation of the module of differentials (#18440) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- Mathlib.lean | 1 + .../Module/Presentation/Differentials.lean | 164 ++++++++++++++++++ 2 files changed, 165 insertions(+) create mode 100644 Mathlib/Algebra/Module/Presentation/Differentials.lean diff --git a/Mathlib.lean b/Mathlib.lean index f3191d04fb72c..589e1a973bc7e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -545,6 +545,7 @@ import Mathlib.Algebra.Module.Pi import Mathlib.Algebra.Module.PointwisePi import Mathlib.Algebra.Module.Presentation.Basic import Mathlib.Algebra.Module.Presentation.Cokernel +import Mathlib.Algebra.Module.Presentation.Differentials import Mathlib.Algebra.Module.Presentation.DirectSum import Mathlib.Algebra.Module.Presentation.Finite import Mathlib.Algebra.Module.Presentation.Free diff --git a/Mathlib/Algebra/Module/Presentation/Differentials.lean b/Mathlib/Algebra/Module/Presentation/Differentials.lean new file mode 100644 index 0000000000000..5ec008c5b1ed6 --- /dev/null +++ b/Mathlib/Algebra/Module/Presentation/Differentials.lean @@ -0,0 +1,164 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import Mathlib.Algebra.Module.Presentation.Basic +import Mathlib.RingTheory.Kaehler.Polynomial +import Mathlib.RingTheory.Kaehler.CotangentComplex +import Mathlib.RingTheory.Presentation + +/-! +# Presentation of the module of differentials + +Given a presentation of a `R`-algebra `S`, we obtain a presentation +of the `S`-module `Ω[S⁄R]`. + +Assume `pres : Algebra.Presentation R S` is a presentation of `S` as an `R`-algebra. +We then have a type `pres.vars` for the generators, a type `pres.rels` for the relations, +and for each `r : pres.rels` we have the relation `pres.relation r` in `pres.Ring` (which is +a ring of polynomials over `R` with variables indexed by `pres.vars`). +Then, `Ω[pres.Ring⁄R]` identifies to the free module on the type `pres.vars`, and +for each `r : pres.rels`, we may consider the image of the differential of `pres.relation r` +in this free module, and by using the (surjective) map `pres.Ring → S`, we obtain +an element `pres.differentialsRelations.relation r` in the free `S`-module on `pres.vars`. +The main result in this file is that `Ω[S⁄R]` identifies to the quotient of the +free `S`-module on `pres.vars` by the submodule generated by these +elements `pres.differentialsRelations.relation r`. We show this as a consequence +of `Algebra.Extension.exact_cotangentComplex_toKaehler` +from the file `Mathlib.RingTheory.Kaehler.CotangentComplex`. + +-/ + +universe w' t w u v + +namespace Algebra.Presentation + +open KaehlerDifferential + +variable {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] + (pres : Algebra.Presentation.{t, w} R S) + +/-- The shape of the presentation by generators and relations of the `S`-module `Ω[S⁄R]` +that is obtained from a presentation of `S` as an `R`-algebra. -/ +@[simps G R] +noncomputable def differentialsRelations : Module.Relations S where + G := pres.vars + R := pres.rels + relation r := + Finsupp.mapRange (algebraMap pres.Ring S) (by simp) + ((mvPolynomialBasis R pres.vars).repr (D _ _ (pres.relation r))) + +namespace differentials + +/-! We obtain here a few compatibilities which allow to compare the two sequences +`(pres.rels →₀ S) → (pres.vars →₀ S) → Ω[S⁄R]` and +`pres.toExtension.Cotangent →ₗ[S] pres.toExtension.CotangentSpace → Ω[S⁄R]`. +Indeed, there is commutative diagram with a surjective map +`hom₁ : (pres.rels →₀ S) →ₗ[S] pres.toExtension.Cotangent` on the left and +bijections on the middle and on the right. Then, the exactness of the first +sequence shall follow from the exactness of the second which is +`Algebra.Extension.exact_cotangentComplex_toKaehler`. -/ + +/-- Same as `comm₂₃` below, but here we have not yet constructed `differentialsSolution`. -/ +lemma comm₂₃' : pres.toExtension.toKaehler.comp pres.cotangentSpaceBasis.repr.symm.toLinearMap = + Finsupp.linearCombination S (fun g ↦ D _ _ (pres.val g)) := by + ext g + dsimp + rw [Basis.repr_symm_apply, Finsupp.linearCombination_single, + Finsupp.linearCombination_single, one_smul, one_smul, + Generators.cotangentSpaceBasis_apply, mapBaseChange_tmul, one_smul] + simp + +/-- The canonical map `(pres.rels →₀ S) →ₗ[S] pres.toExtension.Cotangent`. -/ +noncomputable def hom₁ : (pres.rels →₀ S) →ₗ[S] pres.toExtension.Cotangent := + Finsupp.linearCombination S (fun r ↦ Extension.Cotangent.mk ⟨pres.relation r, by simp⟩) + +lemma hom₁_single (r : pres.rels) : + hom₁ pres (Finsupp.single r 1) = Extension.Cotangent.mk ⟨pres.relation r, by simp⟩ := by + simp [hom₁] + +lemma surjective_hom₁ : Function.Surjective (hom₁ pres) := by + let φ : (pres.rels →₀ S) →ₗ[pres.Ring] pres.toExtension.Cotangent := + { toFun := hom₁ pres + map_add' := by simp + map_smul' := by simp } + change Function.Surjective φ + have h₁ := Algebra.Extension.Cotangent.mk_surjective (P := pres.toExtension) + have h₂ : Submodule.span pres.Ring + (Set.range (fun r ↦ (⟨pres.relation r, by simp⟩ : pres.ker))) = ⊤ := by + refine Submodule.map_injective_of_injective (f := Submodule.subtype pres.ker) + Subtype.coe_injective ?_ + rw [Submodule.map_top, Submodule.range_subtype, Submodule.map_span, + Submodule.coe_subtype, Ideal.submodule_span_eq] + simp only [← pres.span_range_relation_eq_ker] + congr + aesop + rw [← LinearMap.range_eq_top] at h₁ ⊢ + rw [← top_le_iff, ← h₁, LinearMap.range_eq_map, ← h₂] + dsimp + rw [Submodule.map_span_le] + rintro _ ⟨r, rfl⟩ + simp only [LinearMap.mem_range] + refine ⟨Finsupp.single r 1, ?_⟩ + simp only [LinearMap.coe_mk, AddHom.coe_mk, hom₁_single, φ] + rfl + +lemma comm₁₂_single (r : pres.rels) : + pres.toExtension.cotangentComplex (hom₁ pres (Finsupp.single r 1)) = + pres.cotangentSpaceBasis.repr.symm ((differentialsRelations pres).relation r) := by + simp only [hom₁, Finsupp.linearCombination_single, one_smul, differentialsRelations, + Basis.repr_symm_apply, Extension.cotangentComplex_mk] + exact pres.cotangentSpaceBasis.repr.injective (by ext; simp) + +lemma comm₁₂ : pres.toExtension.cotangentComplex.comp (hom₁ pres) = + pres.cotangentSpaceBasis.repr.symm.comp (differentialsRelations pres).map := by + ext r + have := (differentialsRelations pres).map_single + dsimp at this ⊢ + rw [comm₁₂_single, this] + +end differentials + +open differentials in +/-- The `S`-module `Ω[S⁄R]` contains an obvious solution to the system of linear +equations `pres.differentialsRelations.Solution` when `pres` is a presentation +of `S` as an `R`-algebra. -/ +noncomputable def differentialsSolution : + pres.differentialsRelations.Solution (Ω[S⁄R]) where + var g := D _ _ (pres.val g) + linearCombination_var_relation r := by + simp only [differentialsRelations_G, LinearMap.coe_comp, LinearEquiv.coe_coe, + Function.comp_apply, ← comm₂₃', ← comm₁₂_single] + apply DFunLike.congr_fun (Function.Exact.linearMap_comp_eq_zero + (pres.toExtension.exact_cotangentComplex_toKaehler)) + +lemma differentials.comm₂₃ : + pres.toExtension.toKaehler.comp pres.cotangentSpaceBasis.repr.symm.toLinearMap = + pres.differentialsSolution.π := + comm₂₃' pres + +open differentials in +lemma differentialsSolution_isPresentation : + pres.differentialsSolution.IsPresentation := by + rw [Module.Relations.Solution.isPresentation_iff] + constructor + · rw [← Module.Relations.Solution.surjective_π_iff_span_eq_top, ← comm₂₃] + exact Extension.toKaehler_surjective.comp pres.cotangentSpaceBasis.repr.symm.surjective + · rw [← Module.Relations.range_map] + exact Function.Exact.linearMap_ker_eq + ((LinearMap.exact_iff_of_surjective_of_bijective_of_injective + _ _ _ _ (hom₁ pres) + pres.cotangentSpaceBasis.repr.symm.toLinearMap .id + (comm₁₂ pres) (by simpa using comm₂₃ pres) (surjective_hom₁ pres) + (LinearEquiv.bijective _) (Equiv.refl _).injective).2 + pres.toExtension.exact_cotangentComplex_toKaehler) + +/-- The presentation of the `S`-module `Ω[S⁄R]` deduced from a presentation +of `S` as a `R`-algebra. -/ +noncomputable def differentials : Module.Presentation S (Ω[S⁄R]) where + toSolution := differentialsSolution pres + toIsPresentation := pres.differentialsSolution_isPresentation + +end Algebra.Presentation From db422a292b959d8dec7f19c72494ba4915869d78 Mon Sep 17 00:00:00 2001 From: Emily Riehl Date: Wed, 27 Nov 2024 18:16:39 +0000 Subject: [PATCH 360/829] feat(CategoryTheory): coskeletal simplicial objects (#19492) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit A simplicial object `X` is *n-coskeletal* if the identity natural transformation `X` is a right extension of its restriction along `(Truncated.inclusion (n := n)).op` recorded by `rightExtensionInclusion X n`. Co-Authored-By: [Mario Carneiro](https://github.com/digama0) and [Joël Riou](https://github.com/joelriou) --- Mathlib.lean | 3 +- Mathlib/AlgebraicTopology/CechNerve.lean | 2 +- Mathlib/AlgebraicTopology/MooreComplex.lean | 2 +- .../AlgebraicTopology/SimplexCategory.lean | 8 +- .../Basic.lean} | 38 +++++--- .../SimplicialObject/Coskeletal.lean | 91 +++++++++++++++++++ .../SimplicialSet/Basic.lean | 2 +- .../SplitSimplicialObject.lean | 2 +- .../Idempotents/SimplicialObject.lean | 2 +- 9 files changed, 126 insertions(+), 24 deletions(-) rename Mathlib/AlgebraicTopology/{SimplicialObject.lean => SimplicialObject/Basic.lean} (96%) create mode 100644 Mathlib/AlgebraicTopology/SimplicialObject/Coskeletal.lean diff --git a/Mathlib.lean b/Mathlib.lean index 589e1a973bc7e..4b3dcced1bc36 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1042,7 +1042,8 @@ import Mathlib.AlgebraicTopology.Quasicategory.StrictSegal import Mathlib.AlgebraicTopology.SimplexCategory import Mathlib.AlgebraicTopology.SimplicialCategory.Basic import Mathlib.AlgebraicTopology.SimplicialCategory.SimplicialObject -import Mathlib.AlgebraicTopology.SimplicialObject +import Mathlib.AlgebraicTopology.SimplicialObject.Basic +import Mathlib.AlgebraicTopology.SimplicialObject.Coskeletal import Mathlib.AlgebraicTopology.SimplicialSet.Basic import Mathlib.AlgebraicTopology.SimplicialSet.KanComplex import Mathlib.AlgebraicTopology.SimplicialSet.Monoidal diff --git a/Mathlib/AlgebraicTopology/CechNerve.lean b/Mathlib/AlgebraicTopology/CechNerve.lean index 0142017c6b6ab..667421531fc21 100644 --- a/Mathlib/AlgebraicTopology/CechNerve.lean +++ b/Mathlib/AlgebraicTopology/CechNerve.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Adam Topaz. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Adam Topaz -/ -import Mathlib.AlgebraicTopology.SimplicialObject +import Mathlib.AlgebraicTopology.SimplicialObject.Basic import Mathlib.CategoryTheory.Comma.Arrow import Mathlib.CategoryTheory.Limits.Shapes.WidePullbacks import Mathlib.CategoryTheory.Limits.Shapes.FiniteProducts diff --git a/Mathlib/AlgebraicTopology/MooreComplex.lean b/Mathlib/AlgebraicTopology/MooreComplex.lean index 0ce1710a542c8..0276dc0ce398a 100644 --- a/Mathlib/AlgebraicTopology/MooreComplex.lean +++ b/Mathlib/AlgebraicTopology/MooreComplex.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ import Mathlib.Algebra.Homology.HomologicalComplex -import Mathlib.AlgebraicTopology.SimplicialObject +import Mathlib.AlgebraicTopology.SimplicialObject.Basic import Mathlib.CategoryTheory.Abelian.Basic /-! diff --git a/Mathlib/AlgebraicTopology/SimplexCategory.lean b/Mathlib/AlgebraicTopology/SimplexCategory.lean index 7c9ce0278a0c3..115b2bc0a1851 100644 --- a/Mathlib/AlgebraicTopology/SimplexCategory.lean +++ b/Mathlib/AlgebraicTopology/SimplexCategory.lean @@ -725,15 +725,15 @@ instance {n} : Inhabited (Truncated n) := /-- The fully faithful inclusion of the truncated simplex category into the usual simplex category. -/ -def inclusion {n : ℕ} : SimplexCategory.Truncated n ⥤ SimplexCategory := +def inclusion (n : ℕ) : SimplexCategory.Truncated n ⥤ SimplexCategory := fullSubcategoryInclusion _ -instance (n : ℕ) : (inclusion : Truncated n ⥤ _).Full := FullSubcategory.full _ -instance (n : ℕ) : (inclusion : Truncated n ⥤ _).Faithful := FullSubcategory.faithful _ +instance (n : ℕ) : (inclusion n : Truncated n ⥤ _).Full := FullSubcategory.full _ +instance (n : ℕ) : (inclusion n : Truncated n ⥤ _).Faithful := FullSubcategory.faithful _ /-- A proof that the full subcategory inclusion is fully faithful.-/ noncomputable def inclusion.fullyFaithful (n : ℕ) : - (inclusion : Truncated n ⥤ _).op.FullyFaithful := Functor.FullyFaithful.ofFullyFaithful _ + (inclusion n : Truncated n ⥤ _).op.FullyFaithful := Functor.FullyFaithful.ofFullyFaithful _ @[ext] theorem Hom.ext {n} {a b : Truncated n} (f g : a ⟶ b) : diff --git a/Mathlib/AlgebraicTopology/SimplicialObject.lean b/Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean similarity index 96% rename from Mathlib/AlgebraicTopology/SimplicialObject.lean rename to Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean index a7a5cb0b002e5..c7ee1ad189a72 100644 --- a/Mathlib/AlgebraicTopology/SimplicialObject.lean +++ b/Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean @@ -250,7 +250,7 @@ section Truncation /-- The truncation functor from simplicial objects to truncated simplicial objects. -/ def truncation (n : ℕ) : SimplicialObject C ⥤ SimplicialObject.Truncated C n := - (whiskeringLeft _ _ _).obj SimplexCategory.Truncated.inclusion.op + (whiskeringLeft _ _ _).obj (SimplexCategory.Truncated.inclusion n).op end Truncation @@ -259,24 +259,24 @@ noncomputable section /-- The n-skeleton as a functor `SimplicialObject.Truncated C n ⥤ SimplicialObject C`. -/ protected abbrev Truncated.sk (n : ℕ) [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), - SimplexCategory.Truncated.inclusion.op.HasLeftKanExtension F] : + (SimplexCategory.Truncated.inclusion n).op.HasLeftKanExtension F] : SimplicialObject.Truncated C n ⥤ SimplicialObject C := - lan (SimplexCategory.Truncated.inclusion.op) + lan (SimplexCategory.Truncated.inclusion n).op /-- The n-coskeleton as a functor `SimplicialObject.Truncated C n ⥤ SimplicialObject C`. -/ protected abbrev Truncated.cosk (n : ℕ) [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), - SimplexCategory.Truncated.inclusion.op.HasRightKanExtension F] : + (SimplexCategory.Truncated.inclusion n).op.HasRightKanExtension F] : SimplicialObject.Truncated C n ⥤ SimplicialObject C := - ran (SimplexCategory.Truncated.inclusion.op) + ran (SimplexCategory.Truncated.inclusion n).op /-- The n-skeleton as an endofunctor on `SimplicialObject C`. -/ abbrev sk (n : ℕ) [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), - SimplexCategory.Truncated.inclusion.op.HasLeftKanExtension F] : + (SimplexCategory.Truncated.inclusion n).op.HasLeftKanExtension F] : SimplicialObject C ⥤ SimplicialObject C := truncation n ⋙ Truncated.sk n /-- The n-coskeleton as an endofunctor on `SimplicialObject C`. -/ abbrev cosk (n : ℕ) [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), - SimplexCategory.Truncated.inclusion.op.HasRightKanExtension F] : + (SimplexCategory.Truncated.inclusion n).op.HasRightKanExtension F] : SimplicialObject C ⥤ SimplicialObject C := truncation n ⋙ Truncated.cosk n end @@ -288,9 +288,9 @@ respectively define left and right adjoints to `truncation n`.-/ variable (n : ℕ) variable [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), - SimplexCategory.Truncated.inclusion.op.HasRightKanExtension F] + (SimplexCategory.Truncated.inclusion n).op.HasRightKanExtension F] variable [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), - SimplexCategory.Truncated.inclusion.op.HasLeftKanExtension F] + (SimplexCategory.Truncated.inclusion n).op.HasLeftKanExtension F] /-- The adjunction between the n-skeleton and n-truncation.-/ noncomputable def skAdj : Truncated.sk (C := C) n ⊣ truncation n := @@ -300,20 +300,30 @@ noncomputable def skAdj : Truncated.sk (C := C) n ⊣ truncation n := noncomputable def coskAdj : truncation (C := C) n ⊣ Truncated.cosk n := ranAdjunction _ _ +instance : ((sk n).obj X).IsLeftKanExtension ((skAdj n).unit.app _) := by + dsimp [sk, skAdj] + rw [lanAdjunction_unit] + infer_instance + +instance : ((cosk n).obj X).IsRightKanExtension ((coskAdj n).counit.app _) := by + dsimp [cosk, coskAdj] + rw [ranAdjunction_counit] + infer_instance + namespace Truncated /- When the left and right Kan extensions exist and are pointwise Kan extensions, `skAdj n` and `coskAdj n` are respectively coreflective and reflective.-/ variable [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), - SimplexCategory.Truncated.inclusion.op.HasPointwiseRightKanExtension F] + (SimplexCategory.Truncated.inclusion n).op.HasPointwiseRightKanExtension F] variable [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), - SimplexCategory.Truncated.inclusion.op.HasPointwiseLeftKanExtension F] + (SimplexCategory.Truncated.inclusion n).op.HasPointwiseLeftKanExtension F] instance cosk_reflective : IsIso (coskAdj (C := C) n).counit := - reflective' SimplexCategory.Truncated.inclusion.op + reflective' (SimplexCategory.Truncated.inclusion n).op instance sk_coreflective : IsIso (skAdj (C := C) n).unit := - coreflective' SimplexCategory.Truncated.inclusion.op + coreflective' (SimplexCategory.Truncated.inclusion n).op /-- Since `Truncated.inclusion` is fully faithful, so is right Kan extension along it.-/ noncomputable def cosk.fullyFaithful : @@ -675,7 +685,7 @@ section Truncation /-- The truncation functor from cosimplicial objects to truncated cosimplicial objects. -/ def truncation (n : ℕ) : CosimplicialObject C ⥤ CosimplicialObject.Truncated C n := - (whiskeringLeft _ _ _).obj SimplexCategory.Truncated.inclusion + (whiskeringLeft _ _ _).obj (SimplexCategory.Truncated.inclusion n) end Truncation diff --git a/Mathlib/AlgebraicTopology/SimplicialObject/Coskeletal.lean b/Mathlib/AlgebraicTopology/SimplicialObject/Coskeletal.lean new file mode 100644 index 0000000000000..247afdd825990 --- /dev/null +++ b/Mathlib/AlgebraicTopology/SimplicialObject/Coskeletal.lean @@ -0,0 +1,91 @@ +/- +Copyright (c) 2024 Emily Riehl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro, Emily Riehl, Joël Riou +-/ +import Mathlib.AlgebraicTopology.SimplicialObject.Basic +import Mathlib.CategoryTheory.Functor.KanExtension.Adjunction +import Mathlib.CategoryTheory.Functor.KanExtension.Basic + +/-! +# Coskeletal simplicial objects + +The identity natural transformation exhibits a simplicial object `X` as a right extension of its +restriction along `(Truncated.inclusion n).op` recorded by `rightExtensionInclusion X n`. + +The simplicial object `X` is *n-coskeletal* if `rightExtensionInclusion X n` is a right Kan +extension. + +When the ambient category admits right Kan extensions along `(Truncated.inclusion n).op`, +then when `X` is `n`-coskeletal, the unit of `coskAdj n` defines an isomorphism: +`isoCoskOfIsCoskeletal : X ≅ (cosk n).obj X`. + +TODO: Prove that `X` is `n`-coskeletal whenever a certain canonical cone is a limit cone. +-/ + +open Opposite + +open CategoryTheory + +open CategoryTheory.Limits CategoryTheory.Functor SimplexCategory + +universe v u v' u' + +namespace CategoryTheory + +namespace SimplicialObject +variable {C : Type u} [Category.{v} C] +variable (X : SimplicialObject C) (n : ℕ) + +namespace Truncated + +/-- The identity natural transformation exhibits a simplicial set as a right extension of its +restriction along `(Truncated.inclusion n).op`.-/ +@[simps!] +def rightExtensionInclusion : + RightExtension (Truncated.inclusion n).op + ((Truncated.inclusion n).op ⋙ X) := RightExtension.mk _ (𝟙 _) + +end Truncated + +open Truncated + +/-- A simplicial object `X` is `n`-coskeletal when it is the right Kan extension of its restriction +along `(Truncated.inclusion n).op` via the identity natural transformation. -/ +@[mk_iff] +class IsCoskeletal : Prop where + isRightKanExtension : IsRightKanExtension X (𝟙 ((Truncated.inclusion n).op ⋙ X)) + +attribute [instance] IsCoskeletal.isRightKanExtension + +section + +variable [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), + (SimplexCategory.Truncated.inclusion n).op.HasRightKanExtension F] + +/-- If `X` is `n`-coskeletal, then `Truncated.rightExtensionInclusion X n` is a terminal object in +the category `RightExtension (Truncated.inclusion n).op (Truncated.inclusion.op ⋙ X)`. -/ +noncomputable def IsCoskeletal.isUniversalOfIsRightKanExtension [X.IsCoskeletal n] : + (rightExtensionInclusion X n).IsUniversal := by + apply Functor.isUniversalOfIsRightKanExtension + +theorem isCoskeletal_iff_isIso : X.IsCoskeletal n ↔ IsIso ((coskAdj n).unit.app X) := by + rw [isCoskeletal_iff] + exact isRightKanExtension_iff_isIso ((coskAdj n).unit.app X) + ((coskAdj n).counit.app _) (𝟙 _) ((coskAdj n).left_triangle_components X) + +instance [X.IsCoskeletal n] : IsIso ((coskAdj n).unit.app X) := by + rw [← isCoskeletal_iff_isIso] + infer_instance + +/-- The canonical isomorphism `X ≅ (cosk n).obj X` defined when `X` is coskeletal and the +`n`-coskeleton functor exists.-/ +@[simps! hom] +noncomputable def isoCoskOfIsCoskeletal [X.IsCoskeletal n] : X ≅ (cosk n).obj X := + asIso ((coskAdj n).unit.app X) + +end + +end SimplicialObject + +end CategoryTheory diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean index 8c329680dec88..6e545ab1fa48e 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.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: Johan Commelin, Kim Morrison, Adam Topaz -/ -import Mathlib.AlgebraicTopology.SimplicialObject +import Mathlib.AlgebraicTopology.SimplicialObject.Basic import Mathlib.CategoryTheory.Limits.Shapes.Types import Mathlib.CategoryTheory.Yoneda import Mathlib.Data.Fin.VecNotation diff --git a/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean b/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean index db4c3e1b3959c..7e6de96ffecde 100644 --- a/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean +++ b/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Joël Riou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ -import Mathlib.AlgebraicTopology.SimplicialObject +import Mathlib.AlgebraicTopology.SimplicialObject.Basic import Mathlib.CategoryTheory.Limits.Shapes.Products import Mathlib.Data.Fintype.Sigma diff --git a/Mathlib/CategoryTheory/Idempotents/SimplicialObject.lean b/Mathlib/CategoryTheory/Idempotents/SimplicialObject.lean index bcccaf7a24640..0a46ade189089 100644 --- a/Mathlib/CategoryTheory/Idempotents/SimplicialObject.lean +++ b/Mathlib/CategoryTheory/Idempotents/SimplicialObject.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Joël Riou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ -import Mathlib.AlgebraicTopology.SimplicialObject +import Mathlib.AlgebraicTopology.SimplicialObject.Basic import Mathlib.CategoryTheory.Idempotents.FunctorCategories /-! From 0aebe464b32804553b79f0268c4030e3d8219a1c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 27 Nov 2024 19:53:47 +0000 Subject: [PATCH 361/829] refactor(Combinatorics/Additive): generalise Ruzsa's covering lemma to non-abelian groups (#19512) Also state it using a fixed real number and use semantic lemma names. This is much more useful in practice. From GrowthInGroups (LeanCamCombi) --- .../Combinatorics/Additive/RuzsaCovering.lean | 83 ++++++++++--------- 1 file changed, 43 insertions(+), 40 deletions(-) diff --git a/Mathlib/Combinatorics/Additive/RuzsaCovering.lean b/Mathlib/Combinatorics/Additive/RuzsaCovering.lean index 6e7183a4afdf7..825ab42bd37ab 100644 --- a/Mathlib/Combinatorics/Additive/RuzsaCovering.lean +++ b/Mathlib/Combinatorics/Additive/RuzsaCovering.lean @@ -4,68 +4,71 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Algebra.Group.Pointwise.Finset.Basic +import Mathlib.Data.Real.Basic import Mathlib.SetTheory.Cardinal.Finite +import Mathlib.Tactic.Positivity.Finset /-! # Ruzsa's covering lemma -This file proves the Ruzsa covering lemma. This says that, for `s`, `t` finsets, we can cover `s` -with at most `(s + t).card / t.card` copies of `t - t`. - -## TODO - -Merge this file with other prerequisites to Freiman's theorem once we have them. +This file proves the Ruzsa covering lemma. This says that, for `A`, `B` finsets, we can cover `A` +with at most `#(A + B) / #B` copies of `B - B`. -/ +open scoped Pointwise -open Pointwise +variable {G : Type*} [Group G] {K : ℝ} namespace Finset - -variable {α : Type*} [DecidableEq α] [CommGroup α] (s : Finset α) {t : Finset α} +variable [DecidableEq G] {A B : Finset G} /-- **Ruzsa's covering lemma**. -/ @[to_additive "**Ruzsa's covering lemma**"] -theorem exists_subset_mul_div (ht : t.Nonempty) : - ∃ u : Finset α, u.card * t.card ≤ (s * t).card ∧ s ⊆ u * t / t := by - haveI : ∀ u, Decidable ((u : Set α).PairwiseDisjoint (· • t)) := fun u ↦ Classical.dec _ - set C := s.powerset.filter fun u ↦ u.toSet.PairwiseDisjoint (· • t) - obtain ⟨u, hu, hCmax⟩ := C.exists_maximal (filter_nonempty_iff.2 - ⟨∅, empty_mem_powerset _, by rw [coe_empty]; exact Set.pairwiseDisjoint_empty⟩) - rw [mem_filter, mem_powerset] at hu - refine ⟨u, - (card_mul_iff.2 <| pairwiseDisjoint_smul_iff.1 hu.2).ge.trans - (card_le_card <| mul_subset_mul_right hu.1), - fun a ha ↦ ?_⟩ - rw [mul_div_assoc] - by_cases hau : a ∈ u - · exact subset_mul_left _ ht.one_mem_div hau - by_cases H : ∀ b ∈ u, Disjoint (a • t) (b • t) - · refine (hCmax _ ?_ <| ssubset_insert hau).elim +theorem ruzsa_covering_mul (hB : B.Nonempty) (hK : #(A * B) ≤ K * #B) : + ∃ F ⊆ A, #F ≤ K ∧ A ⊆ F * (B / B) := by + haveI : ∀ F, Decidable ((F : Set G).PairwiseDisjoint (· • B)) := fun F ↦ Classical.dec _ + set C := {F ∈ A.powerset | F.toSet.PairwiseDisjoint (· • B)} + obtain ⟨F, hF, hFmax⟩ := C.exists_maximal <| filter_nonempty_iff.2 + ⟨∅, empty_mem_powerset _, by rw [coe_empty]; exact Set.pairwiseDisjoint_empty⟩ + rw [mem_filter, mem_powerset] at hF + obtain ⟨hFA, hF⟩ := hF + refine ⟨F, hFA, le_of_mul_le_mul_right ?_ (by positivity : (0 : ℝ) < #B), fun a ha ↦ ?_⟩ + · calc + (#F * #B : ℝ) = #(F * B) := by + rw [card_mul_iff.2 <| pairwiseDisjoint_smul_iff.1 hF, Nat.cast_mul] + _ ≤ #(A * B) := by gcongr + _ ≤ K * #B := hK + by_cases hau : a ∈ F + · exact subset_mul_left _ hB.one_mem_div hau + by_cases H : ∀ b ∈ F, Disjoint (a • B) (b • B) + · refine (hFmax _ ?_ <| ssubset_insert hau).elim rw [mem_filter, mem_powerset, insert_subset_iff, coe_insert] - exact ⟨⟨ha, hu.1⟩, hu.2.insert fun _ hb _ ↦ H _ hb⟩ + exact ⟨⟨ha, hFA⟩, hF.insert fun _ hb _ ↦ H _ hb⟩ push_neg at H simp_rw [not_disjoint_iff, ← inv_smul_mem_iff] at H obtain ⟨b, hb, c, hc₁, hc₂⟩ := H - refine mem_mul.2 ⟨b, hb, a / b, ?_, by simp⟩ - exact mem_div.2 ⟨_, hc₂, _, hc₁, by simp [inv_mul_eq_div]⟩ + exact mem_mul.2 ⟨b, hb, b⁻¹ * a, mem_div.2 ⟨_, hc₂, _, hc₁, by simp⟩, by simp⟩ + +@[to_additive (attr := deprecated (since := "2024-11-26"))] +alias exists_subset_mul_div := ruzsa_covering_mul end Finset namespace Set -variable {α : Type*} [CommGroup α] {s t : Set α} +variable {A B : Set G} -/-- **Ruzsa's covering lemma** for sets. See also `Finset.exists_subset_mul_div`. -/ -@[to_additive "**Ruzsa's covering lemma**. Version for sets. For finsets, -see `Finset.exists_subset_add_sub`."] -lemma exists_subset_mul_div (hs : s.Finite) (ht' : t.Finite) (ht : t.Nonempty) : - ∃ u : Set α, Nat.card u * Nat.card t ≤ Nat.card (s * t) ∧ s ⊆ u * t / t ∧ u.Finite := by - lift s to Finset α using hs - lift t to Finset α using ht' +/-- **Ruzsa's covering lemma** for sets. See also `Finset.ruzsa_covering_mul`. -/ +@[to_additive "**Ruzsa's covering lemma** for sets. See also `Finset.ruzsa_covering_add`."] +lemma ruzsa_covering_mul (hA : A.Finite) (hB : B.Finite) (hB₀ : B.Nonempty) + (hK : Nat.card (A * B) ≤ K * Nat.card B) : + ∃ F ⊆ A, Nat.card F ≤ K ∧ A ⊆ F * (B / B) ∧ F.Finite := by + lift A to Finset G using hA + lift B to Finset G using hB classical - obtain ⟨u, hu, hsut⟩ := Finset.exists_subset_mul_div s ht - refine ⟨u, ?_⟩ - norm_cast - simp [*] + obtain ⟨F, hFA, hF, hAF⟩ := Finset.ruzsa_covering_mul hB₀ (by simpa [← Finset.coe_mul] using hK) + exact ⟨F, by norm_cast; simp [*]⟩ + +@[to_additive (attr := deprecated (since := "2024-11-26"))] +alias exists_subset_mul_div := ruzsa_covering_mul end Set From 53a5bce08b5099841ffdd06a32c37340fcd2e47a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 27 Nov 2024 20:04:20 +0000 Subject: [PATCH 362/829] chore(Pointwise): move auxiliary lemma out (#19537) This has nothing to do with pointwise actions and doesn't need to be additivised --- .../Algebra/Group/Pointwise/Set/Basic.lean | 30 ------------------- .../Algebra/Group/Pointwise/Set/Finite.lean | 2 +- Mathlib/Order/Monotone/Basic.lean | 24 +++++++++++++++ 3 files changed, 25 insertions(+), 31 deletions(-) diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean index a834f509b4e16..49a613a4045d4 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean @@ -1431,33 +1431,3 @@ variable {ι : Type*} {α : ι → Type*} [∀ i, Inv (α i)] lemma inv_pi (s : Set ι) (t : ∀ i, Set (α i)) : (s.pi t)⁻¹ = s.pi fun i ↦ (t i)⁻¹ := by ext x; simp end Set - -/-! ### Miscellaneous -/ - - -open Set - -open Pointwise - -namespace Group - -@[to_additive] -theorem card_pow_eq_card_pow_card_univ_aux {f : ℕ → ℕ} (h1 : Monotone f) {B : ℕ} (h2 : ∀ n, f n ≤ B) - (h3 : ∀ n, f n = f (n + 1) → f (n + 1) = f (n + 2)) : ∀ k, B ≤ k → f k = f B := by - have key : ∃ n : ℕ, n ≤ B ∧ f n = f (n + 1) := by - contrapose! h2 - suffices ∀ n : ℕ, n ≤ B + 1 → n ≤ f n by exact ⟨B + 1, this (B + 1) (le_refl (B + 1))⟩ - exact fun n => - Nat.rec (fun _ => Nat.zero_le (f 0)) - (fun n ih h => - lt_of_le_of_lt (ih (n.le_succ.trans h)) - (lt_of_le_of_ne (h1 n.le_succ) (h2 n (Nat.succ_le_succ_iff.mp h)))) - n - obtain ⟨n, hn1, hn2⟩ := key - replace key : ∀ k : ℕ, f (n + k) = f (n + k + 1) ∧ f (n + k) = f n := fun k => - Nat.rec ⟨hn2, rfl⟩ (fun k ih => ⟨h3 _ ih.1, ih.1.symm.trans ih.2⟩) k - replace key : ∀ k : ℕ, n ≤ k → f k = f n := fun k hk => - (congr_arg f (Nat.add_sub_of_le hk)).symm.trans (key (k - n)).2 - exact fun k hk => (key k (hn1.trans hk)).trans (key B hn1).symm - -end Group diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Finite.lean b/Mathlib/Algebra/Group/Pointwise/Set/Finite.lean index 77adb76f64340..1f5023b2ef5db 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Finite.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Finite.lean @@ -181,7 +181,7 @@ theorem card_pow_eq_card_pow_card_univ [∀ k : ℕ, DecidablePred (· ∈ S ^ k exact Subtype.ext (mul_right_cancel (Subtype.ext_iff.mp hbc)) have mono : Monotone (fun n ↦ Fintype.card (↥(S ^ n)) : ℕ → ℕ) := monotone_nat_of_le_succ fun n ↦ key a _ _ fun b hb ↦ Set.mul_mem_mul hb ha - refine card_pow_eq_card_pow_card_univ_aux mono (fun n ↦ set_fintype_card_le_univ (S ^ n)) + refine fun _ ↦ Nat.stabilises_of_monotone mono (fun n ↦ set_fintype_card_le_univ (S ^ n)) fun n h ↦ le_antisymm (mono (n + 1).le_succ) (key a⁻¹ (S ^ (n + 2)) (S ^ (n + 1)) ?_) replace h₂ : S ^ n * {a} = S ^ (n + 1) := by have : Fintype (S ^ n * Set.singleton a) := by diff --git a/Mathlib/Order/Monotone/Basic.lean b/Mathlib/Order/Monotone/Basic.lean index c567287e7cdbd..163c8a7c18ccd 100644 --- a/Mathlib/Order/Monotone/Basic.lean +++ b/Mathlib/Order/Monotone/Basic.lean @@ -10,6 +10,7 @@ import Mathlib.Order.Compare import Mathlib.Order.Max import Mathlib.Order.RelClasses import Mathlib.Tactic.Coe +import Mathlib.Tactic.Contrapose import Mathlib.Tactic.Choose /-! @@ -1104,3 +1105,26 @@ alias ⟨Monotone.apply₂, Monotone.of_apply₂⟩ := monotone_iff_apply₂ alias ⟨Antitone.apply₂, Antitone.of_apply₂⟩ := antitone_iff_apply₂ end apply + +/-- A monotone function `f : ℕ → ℕ` bounded by `b`, which is constant after stabilising for the +first time, stabilises in at most `b` steps. -/ +lemma Nat.stabilises_of_monotone {f : ℕ → ℕ} {b n : ℕ} (hfmono : Monotone f) (hfb : ∀ m, f m ≤ b) + (hfstab : ∀ m, f m = f (m + 1) → f (m + 1) = f (m + 2)) (hbn : b ≤ n) : f n = f b := by + obtain ⟨m, hmb, hm⟩ : ∃ m ≤ b, f m = f (m + 1) := by + contrapose! hfb + let rec strictMono : ∀ m ≤ b + 1, m ≤ f m + | 0, _ => Nat.zero_le _ + | m + 1, hmb => (strictMono _ <| m.le_succ.trans hmb).trans_lt <| (hfmono m.le_succ).lt_of_ne <| + hfb _ <| Nat.le_of_succ_le_succ hmb + exact ⟨b + 1, strictMono _ le_rfl⟩ + replace key : ∀ k : ℕ, f (m + k) = f (m + k + 1) ∧ f (m + k) = f m := fun k => + Nat.rec ⟨hm, rfl⟩ (fun k ih => ⟨hfstab _ ih.1, ih.1.symm.trans ih.2⟩) k + replace key : ∀ k ≥ m, f k = f m := fun k hk => + (congr_arg f (Nat.add_sub_of_le hk)).symm.trans (key (k - m)).2 + exact (key n (hmb.trans hbn)).trans (key b hmb).symm + +@[deprecated (since := "2024-11-27")] +alias Group.card_pow_eq_card_pow_card_univ_aux := Nat.stabilises_of_monotone + +@[deprecated (since := "2024-11-27")] +alias Group.card_nsmul_eq_card_nsmulpow_card_univ_aux := Nat.stabilises_of_monotone From 73295cb7b028e8ffafe18f24e14d26ddd4aef72a Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 28 Nov 2024 04:54:35 +0000 Subject: [PATCH 363/829] feat(Normed/Group): add `norm_abs_zsmul` and `norm_neg_one_pow_zsmul` (#19541) --- Mathlib/Analysis/Normed/Group/Basic.lean | 41 ++++++++++++++++++++---- 1 file changed, 34 insertions(+), 7 deletions(-) diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index da9bbaf0e7c48..e0106b8988155 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -47,7 +47,6 @@ normed group variable {𝓕 α ι κ E F G : Type*} open Filter Function Metric Bornology - open ENNReal Filter NNReal Uniformity Pointwise Topology /-- Auxiliary class, endowing a type `E` with a function `norm : E → ℝ` with notation `‖x‖`. This @@ -65,7 +64,6 @@ class NNNorm (E : Type*) where nnnorm : E → ℝ≥0 export Norm (norm) - export NNNorm (nnnorm) @[inherit_doc] @@ -375,6 +373,22 @@ theorem norm_div_rev (a b : E) : ‖a / b‖ = ‖b / a‖ := by @[to_additive (attr := simp) norm_neg] theorem norm_inv' (a : E) : ‖a⁻¹‖ = ‖a‖ := by simpa using norm_div_rev 1 a +@[to_additive (attr := simp) norm_abs_zsmul] +theorem norm_zpow_abs (a : E) (n : ℤ) : ‖a ^ |n|‖ = ‖a ^ n‖ := by + rcases le_total 0 n with hn | hn <;> simp [hn, abs_of_nonneg, abs_of_nonpos] + +@[to_additive (attr := simp) norm_natAbs_smul] +theorem norm_pow_natAbs (a : E) (n : ℤ) : ‖a ^ n.natAbs‖ = ‖a ^ n‖ := by + rw [← zpow_natCast, ← Int.abs_eq_natAbs, norm_zpow_abs] + +@[to_additive norm_isUnit_zsmul] +theorem norm_zpow_isUnit (a : E) {n : ℤ} (hn : IsUnit n) : ‖a ^ n‖ = ‖a‖ := by + rw [← norm_pow_natAbs, Int.isUnit_iff_natAbs_eq.mp hn, pow_one] + +@[simp] +theorem norm_units_zsmul {E : Type*} [SeminormedAddGroup E] (n : ℤˣ) (a : E) : ‖n • a‖ = ‖a‖ := + norm_isUnit_zsmul a n.isUnit + open scoped symmDiff in @[to_additive] theorem dist_mulIndicator (s t : Set α) (f : α → E) (x : α) : @@ -492,7 +506,6 @@ theorem norm_le_norm_add_norm_div (u v : E) : ‖v‖ ≤ ‖u‖ + ‖u / v‖ exact norm_le_norm_add_norm_div' v u alias norm_le_insert' := norm_le_norm_add_norm_sub' - alias norm_le_insert := norm_le_norm_add_norm_sub @[to_additive] @@ -650,8 +663,7 @@ instance (priority := 100) SeminormedGroup.toNNNorm : NNNorm E := ⟨fun a => ⟨‖a‖, norm_nonneg' a⟩⟩ @[to_additive (attr := simp, norm_cast) coe_nnnorm] -theorem coe_nnnorm' (a : E) : (‖a‖₊ : ℝ) = ‖a‖ := - rfl +theorem coe_nnnorm' (a : E) : (‖a‖₊ : ℝ) = ‖a‖ := rfl @[to_additive (attr := simp) coe_comp_nnnorm] theorem coe_comp_nnnorm' : (toReal : ℝ≥0 → ℝ) ∘ (nnnorm : E → ℝ≥0) = norm := @@ -675,8 +687,7 @@ theorem edist_one_right (a : E) : edist a 1 = ‖a‖₊ := by rw [edist_nndist, nndist_one_right] @[to_additive (attr := simp) nnnorm_zero] -theorem nnnorm_one' : ‖(1 : E)‖₊ = 0 := - NNReal.eq norm_one' +theorem nnnorm_one' : ‖(1 : E)‖₊ = 0 := NNReal.eq norm_one' @[to_additive] theorem ne_one_of_nnnorm_ne_zero {a : E} : ‖a‖₊ ≠ 0 → a ≠ 1 := @@ -692,6 +703,22 @@ theorem nnnorm_mul_le' (a b : E) : ‖a * b‖₊ ≤ ‖a‖₊ + ‖b‖₊ := theorem nnnorm_inv' (a : E) : ‖a⁻¹‖₊ = ‖a‖₊ := NNReal.eq <| norm_inv' a +@[to_additive (attr := simp) nnnorm_abs_zsmul] +theorem nnnorm_zpow_abs (a : E) (n : ℤ) : ‖a ^ |n|‖₊ = ‖a ^ n‖₊ := + NNReal.eq <| norm_zpow_abs a n + +@[to_additive (attr := simp) nnnorm_natAbs_smul] +theorem nnnorm_pow_natAbs (a : E) (n : ℤ) : ‖a ^ n.natAbs‖₊ = ‖a ^ n‖₊ := + NNReal.eq <| norm_pow_natAbs a n + +@[to_additive nnnorm_isUnit_zsmul] +theorem nnnorm_zpow_isUnit (a : E) {n : ℤ} (hn : IsUnit n) : ‖a ^ n‖₊ = ‖a‖₊ := + NNReal.eq <| norm_zpow_isUnit a hn + +@[simp] +theorem nnnorm_units_zsmul {E : Type*} [SeminormedAddGroup E] (n : ℤˣ) (a : E) : ‖n • a‖₊ = ‖a‖₊ := + NNReal.eq <| norm_isUnit_zsmul a n.isUnit + @[to_additive (attr := simp)] theorem nndist_one_left (a : E) : nndist 1 a = ‖a‖₊ := by simp [nndist_eq_nnnorm_div] From db0f8f7086c5ce0b254833007d08f9d978f51c22 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Thu, 28 Nov 2024 07:30:47 +0000 Subject: [PATCH 364/829] chore: move unitsEquivNeZero later (#19437) This is only used in a single place, where the later file is already imported. This reduces the imports of a number of low level algebra files. --- Mathlib/Algebra/GroupWithZero/Units/Basic.lean | 15 ++++----------- Mathlib/Algebra/GroupWithZero/Units/Equiv.lean | 16 +++++++++------- 2 files changed, 13 insertions(+), 18 deletions(-) diff --git a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean index 6dcca5a5b9844..7702e037e5cb7 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean @@ -5,12 +5,10 @@ Authors: Johan Commelin -/ import Mathlib.Algebra.Group.Units.Basic import Mathlib.Algebra.GroupWithZero.Basic -import Mathlib.Logic.Equiv.Defs import Mathlib.Tactic.Contrapose import Mathlib.Tactic.Nontriviality import Mathlib.Tactic.Spread import Mathlib.Util.AssertExists -import Mathlib.Data.Subtype /-! # Lemmas about units in a `MonoidWithZero` or a `GroupWithZero`. @@ -20,8 +18,11 @@ We also define `Ring.inverse`, a globally defined function on any ring -/ -- Guard against import creep -assert_not_exists Multiplicative assert_not_exists DenselyOrdered +assert_not_exists Equiv +assert_not_exists Multiplicative +assert_not_exists Subtype.restrict + variable {α M₀ G₀ : Type*} variable [MonoidWithZero M₀] @@ -392,14 +393,6 @@ theorem Ring.inverse_eq_inv (a : G₀) : Ring.inverse a = a⁻¹ := by theorem Ring.inverse_eq_inv' : (Ring.inverse : G₀ → G₀) = Inv.inv := funext Ring.inverse_eq_inv -/-- In a `GroupWithZero` `α`, the unit group `αˣ` is equivalent to the subtype of nonzero -elements. -/ -@[simps] def unitsEquivNeZero [GroupWithZero α] : αˣ ≃ {a : α // a ≠ 0} where - toFun a := ⟨a, a.ne_zero⟩ - invFun a := Units.mk0 _ a.prop - left_inv _ := Units.ext rfl - right_inv _ := rfl - end GroupWithZero section CommGroupWithZero diff --git a/Mathlib/Algebra/GroupWithZero/Units/Equiv.lean b/Mathlib/Algebra/GroupWithZero/Units/Equiv.lean index 9dda8f77c9de9..c93d58a0a6142 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Equiv.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Equiv.lean @@ -12,13 +12,17 @@ import Mathlib.Algebra.GroupWithZero.Units.Basic assert_not_exists DenselyOrdered -variable {G₀ : Type*} +variable {G₀ : Type*} [GroupWithZero G₀] -namespace Equiv - -section GroupWithZero +/-- In a `GroupWithZero` `G₀`, the unit group `G₀ˣ` is equivalent to the subtype of nonzero +elements. -/ +@[simps] def unitsEquivNeZero : G₀ˣ ≃ {a : G₀ // a ≠ 0} where + toFun a := ⟨a, a.ne_zero⟩ + invFun a := Units.mk0 _ a.prop + left_inv _ := Units.ext rfl + right_inv _ := rfl -variable [GroupWithZero G₀] +namespace Equiv /-- Left multiplication by a nonzero element in a `GroupWithZero` is a permutation of the underlying type. -/ @@ -47,6 +51,4 @@ def divRight₀ (a : G₀) (ha : a ≠ 0) : G₀ ≃ G₀ where left_inv _ := by simp [ha] right_inv _ := by simp [ha] -end GroupWithZero - end Equiv From 8d6349409589ff07b4d82f5b59384ba4cf177092 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 28 Nov 2024 08:44:15 +0000 Subject: [PATCH 365/829] feat: locally compact manifolds are finite-dimensional (#19362) Transfer Riesz's theorem to manifolds: locally compact manifolds must be modelled on a finite-dimensional space. One direction already exists: `Manifold.locallyCompact_of_finiteDimensional`, and we provide a converse to it. Co-authored-by: grunweg Co-authored-by: sgouezel Co-authored-by: Michael Rothgang --- .../Manifold/SmoothManifoldWithCorners.lean | 70 +++++++++++++++++-- 1 file changed, 65 insertions(+), 5 deletions(-) diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index 1e7be0aaf95d1..7511ccb7ad482 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -257,13 +257,13 @@ theorem target_eq : I.target = range (I : H → E) := by protected theorem uniqueDiffOn : UniqueDiffOn 𝕜 (range I) := I.target_eq ▸ I.uniqueDiffOn' +@[deprecated (since := "2024-09-30")] +protected alias unique_diff := ModelWithCorners.uniqueDiffOn + theorem range_subset_closure_interior : range I ⊆ closure (interior (range I)) := by rw [← I.target_eq] exact I.target_subset_closure_interior -@[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 @@ -858,6 +858,12 @@ theorem map_extend_nhds {x : M} (hy : x ∈ f.source) : map (f.extend I) (𝓝 x) = 𝓝[range I] f.extend I x := by rwa [extend_coe, comp_apply, ← I.map_nhds_eq, ← f.map_nhds_eq, map_map] +theorem map_extend_nhds_of_mem_interior_range {x : M} (hx : x ∈ f.source) + (h'x : f.extend I x ∈ interior (range I)) : + map (f.extend I) (𝓝 x) = 𝓝 (f.extend I x) := by + rw [f.map_extend_nhds hx, nhdsWithin_eq_nhds] + exact mem_of_superset (isOpen_interior.mem_nhds h'x) interior_subset + theorem map_extend_nhds_of_boundaryless [I.Boundaryless] {x : M} (hx : x ∈ f.source) : map (f.extend I) (𝓝 x) = 𝓝 (f.extend I x) := by rw [f.map_extend_nhds hx, I.range_eq_univ, nhdsWithin_univ] @@ -872,6 +878,12 @@ theorem extend_image_nhd_mem_nhds_of_boundaryless [I.Boundaryless] {x} (hx : x rw [← f.map_extend_nhds_of_boundaryless hx, Filter.mem_map] filter_upwards [h] using subset_preimage_image (f.extend I) s +theorem extend_image_nhd_mem_nhds_of_mem_interior_range {x} (hx : x ∈ f.source) + (h'x : f.extend I x ∈ interior (range I)) {s : Set M} (h : s ∈ 𝓝 x) : + (f.extend I) '' s ∈ 𝓝 ((f.extend I) x) := by + rw [← f.map_extend_nhds_of_mem_interior_range hx h'x, Filter.mem_map] + filter_upwards [h] using subset_preimage_image (f.extend I) s + theorem extend_target_subset_range : (f.extend I).target ⊆ range I := by simp only [mfld_simps] lemma interior_extend_target_subset_interior_range : @@ -1169,6 +1181,12 @@ theorem map_extChartAt_nhds_of_boundaryless [I.Boundaryless] (x : M) : rw [extChartAt] exact map_extend_nhds_of_boundaryless (chartAt H x) (mem_chart_source H x) +theorem extChartAt_image_nhd_mem_nhds_of_mem_interior_range {x y} (hx : y ∈ (extChartAt I x).source) + (h'x : extChartAt I x y ∈ interior (range I)) {s : Set M} (h : s ∈ 𝓝 y) : + (extChartAt I x) '' s ∈ 𝓝 (extChartAt I x y) := by + rw [extChartAt] + exact extend_image_nhd_mem_nhds_of_mem_interior_range _ (by simpa using hx) h'x h + variable {x} in theorem extChartAt_image_nhd_mem_nhds_of_boundaryless [I.Boundaryless] {x : M} (hx : s ∈ 𝓝 x) : extChartAt I x '' s ∈ 𝓝 (extChartAt I x x) := by @@ -1189,11 +1207,15 @@ theorem extChartAt_target_mem_nhdsWithin_of_mem {x : M} {y : E} (hy : y ∈ (ext apply extChartAt_target_mem_nhdsWithin' exact (extChartAt I x).map_target hy -theorem extChartAt_target_union_comp_range_mem_nhds_of_mem {y : E} {x : M} +theorem extChartAt_target_union_compl_range_mem_nhds_of_mem {y : E} {x : M} (hy : y ∈ (extChartAt I x).target) : (extChartAt I x).target ∪ (range I)ᶜ ∈ 𝓝 y := by rw [← nhdsWithin_univ, ← union_compl_self (range I), nhdsWithin_union] exact Filter.union_mem_sup (extChartAt_target_mem_nhdsWithin_of_mem hy) self_mem_nhdsWithin +@[deprecated (since := "2024-11-27")] alias +extChartAt_target_union_comp_range_mem_nhds_of_mem := +extChartAt_target_union_compl_range_mem_nhds_of_mem + /-- If we're boundaryless, `extChartAt` has open target -/ theorem isOpen_extChartAt_target [I.Boundaryless] (x : M) : IsOpen (extChartAt I x).target := by simp_rw [extChartAt_target, I.range_eq_univ, inter_univ] @@ -1270,7 +1292,7 @@ lemma extChartAt_target_subset_closure_interior {x : M} : rw [mem_closure_iff_nhds] intro t ht have A : t ∩ ((extChartAt I x).target ∪ (range I)ᶜ) ∈ 𝓝 y := - inter_mem ht (extChartAt_target_union_comp_range_mem_nhds_of_mem hy) + inter_mem ht (extChartAt_target_union_compl_range_mem_nhds_of_mem hy) have B : y ∈ closure (interior (range I)) := by apply I.range_subset_closure_interior (extChartAt_target_subset_range x hy) obtain ⟨z, ⟨tz, h'z⟩, hz⟩ : @@ -1280,6 +1302,13 @@ lemma extChartAt_target_subset_closure_interior {x : M} : have h''z : z ∈ (extChartAt I x).target := by simpa [interior_subset hz] using h'z exact (extChartAt_target_eventuallyEq_of_mem h''z).symm.mem_interior hz +variable (I) in +theorem interior_extChartAt_target_nonempty (x : M) : + (interior (extChartAt I x).target).Nonempty := by + by_contra! H + have := extChartAt_target_subset_closure_interior (mem_extChartAt_target (I := I) x) + simp only [H, closure_empty, mem_empty_iff_false] at this + lemma extChartAt_mem_closure_interior {x₀ x : M} (hx : x ∈ closure (interior s)) (h'x : x ∈ (extChartAt I x₀).source) : extChartAt I x₀ x ∈ @@ -1482,6 +1511,7 @@ theorem writtenInExtChartAt_chartAt_symm_comp [ChartedSpace H H'] (x : M') {y} end ExtendedCharts section Topology + -- Let `M` be a topological manifold over the field 𝕜. variable {E : Type*} {𝕜 : Type*} [NontriviallyNormedField 𝕜] @@ -1497,6 +1527,36 @@ lemma Manifold.locallyCompact_of_finiteDimensional have : LocallyCompactSpace H := I.locallyCompactSpace exact ChartedSpace.locallyCompactSpace H M +variable (M) + +/-- A locally compact manifold must be modelled on a locally compact space. -/ +lemma LocallyCompactSpace.of_locallyCompact_manifold (I : ModelWithCorners 𝕜 E H) + [h : Nonempty M] [LocallyCompactSpace M] : + LocallyCompactSpace E := by + rcases h with ⟨x⟩ + obtain ⟨y, hy⟩ := interior_extChartAt_target_nonempty I x + have h'y : y ∈ (extChartAt I x).target := interior_subset hy + obtain ⟨s, hmem, hss, hcom⟩ := + LocallyCompactSpace.local_compact_nhds ((extChartAt I x).symm y) (extChartAt I x).source + ((isOpen_extChartAt_source x).mem_nhds ((extChartAt I x).map_target h'y)) + have : IsCompact <| (extChartAt I x) '' s := + hcom.image_of_continuousOn <| (continuousOn_extChartAt x).mono hss + apply this.locallyCompactSpace_of_mem_nhds_of_addGroup (x := y) + rw [← (extChartAt I x).right_inv h'y] + apply extChartAt_image_nhd_mem_nhds_of_mem_interior_range + (PartialEquiv.map_target (extChartAt I x) h'y) _ hmem + simp only [(extChartAt I x).right_inv h'y] + exact interior_mono (extChartAt_target_subset_range x) hy + +/-- Riesz's theorem applied to manifolds: a locally compact manifolds must be modelled on a + finite-dimensional space. This is the converse to + `Manifold.locallyCompact_of_finiteDimensional`. -/ +theorem FiniteDimensional.of_locallyCompact_manifold + [CompleteSpace 𝕜] (I : ModelWithCorners 𝕜 E H) [Inhabited M] [LocallyCompactSpace M] : + FiniteDimensional 𝕜 E := by + have := LocallyCompactSpace.of_locallyCompact_manifold M I + exact FiniteDimensional.of_locallyCompactSpace 𝕜 + end Topology section TangentSpace From 1109970b9069835cb9d4cd69bf196c9fff7516f4 Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Thu, 28 Nov 2024 12:19:19 +0000 Subject: [PATCH 366/829] feat: `IntermediateField.rank_sup_le` (#19527) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit New results: In `Mathlib/RingTheory/AlgebraicIndependent.lean`: - `IsTranscendenceBasis.lift_(cardinalMk|rank)_eq_max_lift`: if `x` is a transcendence basis of `E/F` and is not empty, then `[E:F] = #E = max(#F, #x, ℵ₀)` - `Algebra.Transcendental.rank_eq_cardinalMk`: if `E/F` is transcendental, then `[E:F] = #E`. (a corollary of above result) - `IntermediateField.rank_sup_le`: if `A` and `B` are intermediate fields of `E/F`, then `[AB:F] ≤ [A:F] * [B:F]` (an application of above result) In `Mathlib/RingTheory/Algebraic.lean`: - `[Algebra.]Transcendental.infinite`: a ring has infinitely many elements if it has a transcendental element. In `Mathlib/FieldTheory/MvRatFunc/Rank.lean` (new file): - `MvRatFunc.rank_eq_max_lift`: rank of multivariate rational function field. Note that we don't have `MvRatFunc` yet, so the statement uses `FractionRing (MvPolynomial ...)`. Eventually the theory of `MvRatFunc` should go this directory. Also move the results `Algebra.IsAlgebraic.[lift_]cardinalMk_le_XXX` from `Mathlib/FieldTheory/IsAlgClosed/Classification.lean` to a new file `Mathlib/RingTheory/Algebraic/Cardinality.lean` which has fewer imports. --- Mathlib.lean | 2 + .../IsAlgClosed/Classification.lean | 54 +------------ Mathlib/FieldTheory/MvRatFunc/Rank.lean | 35 +++++++++ Mathlib/RingTheory/Algebraic.lean | 13 ++++ Mathlib/RingTheory/Algebraic/Cardinality.lean | 75 +++++++++++++++++++ Mathlib/RingTheory/AlgebraicIndependent.lean | 63 +++++++++++++++- 6 files changed, 188 insertions(+), 54 deletions(-) create mode 100644 Mathlib/FieldTheory/MvRatFunc/Rank.lean create mode 100644 Mathlib/RingTheory/Algebraic/Cardinality.lean diff --git a/Mathlib.lean b/Mathlib.lean index 4b3dcced1bc36..2ac571937ef73 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2985,6 +2985,7 @@ import Mathlib.FieldTheory.Minpoly.Field import Mathlib.FieldTheory.Minpoly.IsConjRoot import Mathlib.FieldTheory.Minpoly.IsIntegrallyClosed import Mathlib.FieldTheory.Minpoly.MinpolyDiv +import Mathlib.FieldTheory.MvRatFunc.Rank import Mathlib.FieldTheory.Normal import Mathlib.FieldTheory.NormalClosure import Mathlib.FieldTheory.Perfect @@ -4198,6 +4199,7 @@ import Mathlib.RingTheory.Adjoin.Tower import Mathlib.RingTheory.AdjoinRoot import Mathlib.RingTheory.AlgebraTower import Mathlib.RingTheory.Algebraic +import Mathlib.RingTheory.Algebraic.Cardinality import Mathlib.RingTheory.AlgebraicIndependent import Mathlib.RingTheory.Artinian import Mathlib.RingTheory.Bezout diff --git a/Mathlib/FieldTheory/IsAlgClosed/Classification.lean b/Mathlib/FieldTheory/IsAlgClosed/Classification.lean index c7b423a2f784e..7cd6d8c7278e5 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/Classification.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/Classification.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.Algebra.ZMod import Mathlib.Algebra.MvPolynomial.Cardinal import Mathlib.Algebra.Polynomial.Cardinal import Mathlib.FieldTheory.IsAlgClosed.Basic +import Mathlib.RingTheory.Algebraic.Cardinality import Mathlib.RingTheory.AlgebraicIndependent /-! @@ -29,59 +30,6 @@ open scoped Cardinal Polynomial open Cardinal -section AlgebraicClosure - -namespace Algebra.IsAlgebraic - -variable (R L : Type u) [CommRing R] [CommRing L] [IsDomain L] [Algebra R L] -variable [NoZeroSMulDivisors R L] [Algebra.IsAlgebraic R L] - -theorem cardinalMk_le_sigma_polynomial : - #L ≤ #(Σ p : R[X], { x : L // x ∈ p.aroots L }) := - @mk_le_of_injective L (Σ p : R[X], {x : L | x ∈ p.aroots L}) - (fun x : L => - let p := Classical.indefiniteDescription _ (Algebra.IsAlgebraic.isAlgebraic x) - ⟨p.1, x, by - dsimp - have h : p.1.map (algebraMap R L) ≠ 0 := by - rw [Ne, ← Polynomial.degree_eq_bot, - Polynomial.degree_map_eq_of_injective (NoZeroSMulDivisors.algebraMap_injective R L), - Polynomial.degree_eq_bot] - exact p.2.1 - rw [Polynomial.mem_roots h, Polynomial.IsRoot, Polynomial.eval_map, ← Polynomial.aeval_def, - p.2.2]⟩) - fun x y => by - intro h - simp? at h says simp only [Set.coe_setOf, ne_eq, Set.mem_setOf_eq, Sigma.mk.inj_iff] at h - refine (Subtype.heq_iff_coe_eq ?_).1 h.2 - simp only [h.1, forall_true_iff] - -@[deprecated (since := "2024-11-10")] -alias cardinal_mk_le_sigma_polynomial := cardinalMk_le_sigma_polynomial - -/-- The cardinality of an algebraic extension is at most the maximum of the cardinality -of the base ring or `ℵ₀`. -/ -@[stacks 09GK] -theorem cardinalMk_le_max : #L ≤ max #R ℵ₀ := - calc - #L ≤ #(Σ p : R[X], { x : L // x ∈ p.aroots L }) := - cardinalMk_le_sigma_polynomial R L - _ = Cardinal.sum fun p : R[X] => #{x : L | x ∈ p.aroots L} := by - rw [← mk_sigma]; rfl - _ ≤ Cardinal.sum.{u, u} fun _ : R[X] => ℵ₀ := - (sum_le_sum _ _ fun _ => (Multiset.finite_toSet _).lt_aleph0.le) - _ = #(R[X]) * ℵ₀ := sum_const' _ _ - _ ≤ max (max #(R[X]) ℵ₀) ℵ₀ := mul_le_max _ _ - _ ≤ max (max (max #R ℵ₀) ℵ₀) ℵ₀ := - (max_le_max (max_le_max Polynomial.cardinalMk_le_max le_rfl) le_rfl) - _ = max #R ℵ₀ := by simp only [max_assoc, max_comm ℵ₀, max_left_comm ℵ₀, max_self] - -@[deprecated (since := "2024-11-10")] alias cardinal_mk_le_max := cardinalMk_le_max - -end Algebra.IsAlgebraic - -end AlgebraicClosure - namespace IsAlgClosed section Classification diff --git a/Mathlib/FieldTheory/MvRatFunc/Rank.lean b/Mathlib/FieldTheory/MvRatFunc/Rank.lean new file mode 100644 index 0000000000000..e6f6bcac0babe --- /dev/null +++ b/Mathlib/FieldTheory/MvRatFunc/Rank.lean @@ -0,0 +1,35 @@ +/- +Copyright (c) 2024 Jz Pan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jz Pan +-/ +import Mathlib.Algebra.MvPolynomial.Cardinal +import Mathlib.RingTheory.Algebraic +import Mathlib.RingTheory.Localization.Cardinality +import Mathlib.RingTheory.MvPolynomial + +/-! +# Rank of multivariate rational function field +-/ + +noncomputable section + +universe u v + +open Cardinal in +theorem MvRatFunc.rank_eq_max_lift + {σ : Type u} {F : Type v} [Field F] [Nonempty σ] : + Module.rank F (FractionRing (MvPolynomial σ F)) = lift.{u} #F ⊔ lift.{v} #σ ⊔ ℵ₀ := by + let R := MvPolynomial σ F + let K := FractionRing R + refine ((rank_le_card _ _).trans ?_).antisymm ?_ + · rw [FractionRing.cardinalMk, MvPolynomial.cardinalMk_eq_max_lift] + have hinj := IsFractionRing.injective R K + have h1 := (IsScalarTower.toAlgHom F R K).toLinearMap.rank_le_of_injective hinj + rw [MvPolynomial.rank_eq_lift, mk_finsupp_nat, lift_max, lift_aleph0, max_le_iff] at h1 + obtain ⟨i⟩ := ‹Nonempty σ› + have hx : Transcendental F (algebraMap R K (MvPolynomial.X i)) := + (transcendental_algebraMap_iff hinj).2 (MvPolynomial.transcendental_X F i) + have h2 := hx.linearIndependent_sub_inv.cardinal_lift_le_rank + rw [lift_id'.{v, u}, lift_umax.{v, u}] at h2 + exact max_le (max_le h2 h1.1) h1.2 diff --git a/Mathlib/RingTheory/Algebraic.lean b/Mathlib/RingTheory/Algebraic.lean index 4752101e3862d..deafa7a49ce6a 100644 --- a/Mathlib/RingTheory/Algebraic.lean +++ b/Mathlib/RingTheory/Algebraic.lean @@ -951,3 +951,16 @@ theorem transcendental_supported_X_iff [Nontrivial R] {i : σ} {s : Set σ} : transcendental_supported_polynomial_aeval_X_iff R (i := i) (s := s) (f := Polynomial.X) end MvPolynomial + +section Infinite + +theorem Transcendental.infinite {R A : Type*} [CommRing R] [Ring A] [Algebra R A] + [Nontrivial R] {x : A} (hx : Transcendental R x) : Infinite A := + .of_injective _ (transcendental_iff_injective.mp hx) + +theorem Algebra.Transcendental.infinite (R A : Type*) [CommRing R] [Ring A] [Algebra R A] + [Nontrivial R] [Algebra.Transcendental R A] : Infinite A := + have ⟨x, hx⟩ := ‹Algebra.Transcendental R A› + hx.infinite + +end Infinite diff --git a/Mathlib/RingTheory/Algebraic/Cardinality.lean b/Mathlib/RingTheory/Algebraic/Cardinality.lean new file mode 100644 index 0000000000000..43311bc18a58e --- /dev/null +++ b/Mathlib/RingTheory/Algebraic/Cardinality.lean @@ -0,0 +1,75 @@ +/- +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.RingTheory.Algebraic + +/-! +# Cardinality of algebraic extensions + +This file contains results on cardinality of algebraic extensions. +-/ + + +universe u v + +open scoped Cardinal Polynomial + +open Cardinal + +namespace Algebra.IsAlgebraic + +variable (R : Type u) [CommRing R] (L : Type v) [CommRing L] [IsDomain L] [Algebra R L] +variable [NoZeroSMulDivisors R L] [Algebra.IsAlgebraic R L] + +theorem lift_cardinalMk_le_sigma_polynomial : + lift.{u} #L ≤ #(Σ p : R[X], { x : L // x ∈ p.aroots L }) := by + have := @lift_mk_le_lift_mk_of_injective L (Σ p : R[X], {x : L | x ∈ p.aroots L}) + (fun x : L => + let p := Classical.indefiniteDescription _ (Algebra.IsAlgebraic.isAlgebraic x) + ⟨p.1, x, by + dsimp + have := (Polynomial.map_ne_zero_iff (NoZeroSMulDivisors.algebraMap_injective R L)).2 p.2.1 + rw [Polynomial.mem_roots this, Polynomial.IsRoot, Polynomial.eval_map, + ← Polynomial.aeval_def, p.2.2]⟩) + fun x y => by + intro h + simp only [Set.coe_setOf, ne_eq, Set.mem_setOf_eq, Sigma.mk.inj_iff] at h + refine (Subtype.heq_iff_coe_eq ?_).1 h.2 + simp only [h.1, forall_true_iff] + rwa [lift_umax, lift_id'.{v}] at this + +theorem lift_cardinalMk_le_max : lift.{u} #L ≤ lift.{v} #R ⊔ ℵ₀ := + calc + lift.{u} #L ≤ #(Σ p : R[X], { x : L // x ∈ p.aroots L }) := + lift_cardinalMk_le_sigma_polynomial R L + _ = Cardinal.sum fun p : R[X] => #{x : L | x ∈ p.aroots L} := by + rw [← mk_sigma]; rfl + _ ≤ Cardinal.sum.{u, v} fun _ : R[X] => ℵ₀ := + (sum_le_sum _ _ fun _ => (Multiset.finite_toSet _).lt_aleph0.le) + _ = lift.{v} #(R[X]) * ℵ₀ := by rw [sum_const, lift_aleph0] + _ ≤ lift.{v} (#R ⊔ ℵ₀) ⊔ ℵ₀ ⊔ ℵ₀ := (mul_le_max _ _).trans <| by + gcongr; simp only [lift_le, Polynomial.cardinalMk_le_max] + _ = _ := by simp + +variable (L : Type u) [CommRing L] [IsDomain L] [Algebra R L] +variable [NoZeroSMulDivisors R L] [Algebra.IsAlgebraic R L] + +theorem cardinalMk_le_sigma_polynomial : + #L ≤ #(Σ p : R[X], { x : L // x ∈ p.aroots L }) := by + simpa only [lift_id] using lift_cardinalMk_le_sigma_polynomial R L + +@[deprecated (since := "2024-11-10")] +alias cardinal_mk_le_sigma_polynomial := cardinalMk_le_sigma_polynomial + +/-- The cardinality of an algebraic extension is at most the maximum of the cardinality +of the base ring or `ℵ₀`. -/ +@[stacks 09GK] +theorem cardinalMk_le_max : #L ≤ max #R ℵ₀ := by + simpa only [lift_id] using lift_cardinalMk_le_max R L + +@[deprecated (since := "2024-11-10")] alias cardinal_mk_le_max := cardinalMk_le_max + +end Algebra.IsAlgebraic diff --git a/Mathlib/RingTheory/AlgebraicIndependent.lean b/Mathlib/RingTheory/AlgebraicIndependent.lean index 7109a5bc4b402..a82ccfeeb3732 100644 --- a/Mathlib/RingTheory/AlgebraicIndependent.lean +++ b/Mathlib/RingTheory/AlgebraicIndependent.lean @@ -6,6 +6,8 @@ Authors: Chris Hughes import Mathlib.Algebra.MvPolynomial.Monad import Mathlib.Data.Fin.Tuple.Reflection import Mathlib.FieldTheory.Adjoin +import Mathlib.FieldTheory.MvRatFunc.Rank +import Mathlib.RingTheory.Algebraic.Cardinality import Mathlib.RingTheory.MvPolynomial.Basic /-! @@ -56,7 +58,7 @@ open Function Set Subalgebra MvPolynomial Algebra open scoped Classical -universe x u v w +universe u v w variable {ι : Type*} {ι' : Type*} (R : Type*) {K : Type*} variable {A : Type*} {A' : Type*} @@ -741,3 +743,62 @@ theorem algebraicIndependent_empty [Nontrivial A] : algebraicIndependent_empty_type end Field + +section RankAndCardinality + +open Cardinal + +theorem IsTranscendenceBasis.lift_cardinalMk_eq_max_lift + {F : Type u} {E : Type v} [CommRing F] [Nontrivial F] [CommRing E] [IsDomain E] [Algebra F E] + {ι : Type w} {x : ι → E} [Nonempty ι] (hx : IsTranscendenceBasis F x) : + lift.{max u w} #E = lift.{max v w} #F ⊔ lift.{max u v} #ι ⊔ ℵ₀ := by + let K := Algebra.adjoin F (Set.range x) + suffices #E = #K by simp [this, ← lift_mk_eq'.2 ⟨hx.1.aevalEquiv.toEquiv⟩] + haveI : Algebra.IsAlgebraic K E := hx.isAlgebraic + refine le_antisymm ?_ (mk_le_of_injective Subtype.val_injective) + haveI : Infinite K := hx.1.aevalEquiv.infinite_iff.1 inferInstance + simpa only [sup_eq_left.2 (aleph0_le_mk K)] using Algebra.IsAlgebraic.cardinalMk_le_max K E + +theorem IsTranscendenceBasis.lift_rank_eq_max_lift + {F : Type u} {E : Type v} [Field F] [Field E] [Algebra F E] + {ι : Type w} {x : ι → E} [Nonempty ι] (hx : IsTranscendenceBasis F x) : + lift.{max u w} (Module.rank F E) = lift.{max v w} #F ⊔ lift.{max u v} #ι ⊔ ℵ₀ := by + let K := IntermediateField.adjoin F (Set.range x) + haveI : Algebra.IsAlgebraic K E := hx.isAlgebraic_field + rw [← rank_mul_rank F K E, lift_mul, ← hx.1.aevalEquivField.toLinearEquiv.lift_rank_eq, + MvRatFunc.rank_eq_max_lift, lift_max, lift_max, lift_lift, lift_lift, lift_aleph0] + refine mul_eq_left le_sup_right ((lift_le.2 ((rank_le_card K E).trans + (Algebra.IsAlgebraic.cardinalMk_le_max K E))).trans_eq ?_) (by simp [rank_pos.ne']) + simp [← lift_mk_eq'.2 ⟨hx.1.aevalEquivField.toEquiv⟩] + +theorem Algebra.Transcendental.rank_eq_cardinalMk + (F : Type u) (E : Type v) [Field F] [Field E] [Algebra F E] [Algebra.Transcendental F E] : + Module.rank F E = #E := by + obtain ⟨ι, x, hx⟩ := exists_isTranscendenceBasis' _ (algebraMap F E).injective + haveI := hx.nonempty_iff_transcendental.2 ‹_› + simpa [← hx.lift_cardinalMk_eq_max_lift] using hx.lift_rank_eq_max_lift + +theorem IntermediateField.rank_sup_le + {F : Type u} {E : Type v} [Field F] [Field E] [Algebra F E] (A B : IntermediateField F E) : + Module.rank F ↥(A ⊔ B) ≤ Module.rank F A * Module.rank F B := by + by_cases hA : Algebra.IsAlgebraic F A + · exact rank_sup_le_of_isAlgebraic A B (Or.inl hA) + by_cases hB : Algebra.IsAlgebraic F B + · exact rank_sup_le_of_isAlgebraic A B (Or.inr hB) + rw [← Algebra.transcendental_iff_not_isAlgebraic] at hA hB + haveI : Algebra.Transcendental F ↥(A ⊔ B) := .ringHom_of_comp_eq (RingHom.id F) + (inclusion le_sup_left) Function.surjective_id (inclusion_injective _) rfl + haveI := Algebra.Transcendental.infinite F A + haveI := Algebra.Transcendental.infinite F B + simp_rw [Algebra.Transcendental.rank_eq_cardinalMk] + rw [sup_def, mul_mk_eq_max, ← Cardinal.lift_le.{u}] + refine (lift_cardinalMk_adjoin_le _ _).trans ?_ + calc + _ ≤ Cardinal.lift.{v} #F ⊔ Cardinal.lift.{u} (#A ⊔ #B) ⊔ ℵ₀ := by + gcongr + rw [Cardinal.lift_le] + exact (mk_union_le _ _).trans_eq (by simp) + _ = _ := by + simp [lift_mk_le_lift_mk_of_injective (algebraMap F A).injective] + +end RankAndCardinality From 4b33b00f101ad37a30eb73b2e3e1a05101f8fee8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Martin=20Dvo=C5=99=C3=A1k?= Date: Thu, 28 Nov 2024 13:23:50 +0000 Subject: [PATCH 367/829] feat: adjoining identity matrices to TU matrices (#19076) Co-authored-by: Bhavik Mehta Co-authored-by: Eric Wieser --- .../Matrix/Determinant/TotallyUnimodular.lean | 150 ++++++++++++------ 1 file changed, 105 insertions(+), 45 deletions(-) diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean index 03da13b5d1757..8300120633c79 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2024 Martin Dvorak. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Martin Dvorak, Vladimir Kolmogorov, Ivan Sergeev +Authors: Martin Dvorak, Vladimir Kolmogorov, Ivan Sergeev, Bhavik Mehta -/ import Mathlib.LinearAlgebra.Matrix.Determinant.Basic import Mathlib.Data.Matrix.ColumnRowPartitioned @@ -27,7 +27,7 @@ This file defines totally unimodular matrices and provides basic API for them. namespace Matrix -variable {m n R : Type*} [CommRing R] +variable {m m' n n' R : Type*} [CommRing R] /-- `A.IsTotallyUnimodular` means that every square submatrix of `A` (not necessarily contiguous) has determinant `0` or `1` or `-1`; that is, the determinant is in the range of `SignType.cast`. -/ @@ -40,12 +40,12 @@ lemma isTotallyUnimodular_iff (A : Matrix m n R) : A.IsTotallyUnimodular ↔ (A.submatrix f g).det ∈ Set.range SignType.cast := by constructor <;> intro hA · intro k f g - by_cases h : f.Injective ∧ g.Injective - · exact hA k f g h.1 h.2 - · refine ⟨0, ?_⟩ + by_cases hfg : f.Injective ∧ g.Injective + · exact hA k f g hfg.1 hfg.2 + · use 0 rw [SignType.coe_zero, eq_comm] - simp_rw [not_and_or, Function.not_injective_iff] at h - obtain ⟨i, j, hfij, hij⟩ | ⟨i, j, hgij, hij⟩ := h + simp_rw [not_and_or, Function.not_injective_iff] at hfg + obtain ⟨i, j, hfij, hij⟩ | ⟨i, j, hgij, hij⟩ := hfg · rw [← det_transpose, transpose_submatrix] apply det_zero_of_column_eq hij.symm simp [hfij] @@ -66,16 +66,13 @@ lemma isTotallyUnimodular_iff_fintype.{w} (A : Matrix m n R) : A.IsTotallyUnimod specialize hA (ULift (Fin k)) (f ∘ Equiv.ulift) (g ∘ Equiv.ulift) rwa [←submatrix_submatrix, det_submatrix_equiv_self] at hA -lemma IsTotallyUnimodular.apply {A : Matrix m n R} (hA : A.IsTotallyUnimodular) - (i : m) (j : n) : +lemma IsTotallyUnimodular.apply {A : Matrix m n R} (hA : A.IsTotallyUnimodular) (i : m) (j : n) : A i j ∈ Set.range SignType.cast := by - let f : Fin 1 → m := (fun _ => i) - let g : Fin 1 → n := (fun _ => j) - convert hA 1 f g (Function.injective_of_subsingleton f) (Function.injective_of_subsingleton g) - exact (det_fin_one (A.submatrix f g)).symm + rw [isTotallyUnimodular_iff] at hA + simpa using hA 1 (fun _ => i) (fun _ => j) -lemma IsTotallyUnimodular.submatrix {A : Matrix m n R} (hA : A.IsTotallyUnimodular) {k : ℕ} - (f : Fin k → m) (g : Fin k → n) : +lemma IsTotallyUnimodular.submatrix {A : Matrix m n R} (f : m' → m) (g : n' → n) + (hA : A.IsTotallyUnimodular) : (A.submatrix f g).IsTotallyUnimodular := by simp only [isTotallyUnimodular_iff, submatrix_submatrix] at hA ⊢ intro _ _ _ @@ -91,38 +88,101 @@ lemma transpose_isTotallyUnimodular_iff (A : Matrix m n R) : Aᵀ.IsTotallyUnimodular ↔ A.IsTotallyUnimodular := by constructor <;> apply IsTotallyUnimodular.transpose -lemma mapEquiv_isTotallyUnimodular {X' Y' : Type*} (A : Matrix m n R) (eX : X' ≃ m) (eY : Y' ≃ n) : - IsTotallyUnimodular ((A · ∘ eY) ∘ eX) ↔ A.IsTotallyUnimodular := by - rw [isTotallyUnimodular_iff, isTotallyUnimodular_iff] - constructor <;> intro hA k f g - · simpa [submatrix] using hA k (eX.symm ∘ f) (eY.symm ∘ g) - · simpa [submatrix] using hA k (eX ∘ f) (eY ∘ g) - -lemma fromRows_row0_isTotallyUnimodular_iff (A : Matrix m n R) {m' : Type*} : - (fromRows A (row m' 0)).IsTotallyUnimodular ↔ A.IsTotallyUnimodular := by - rw [isTotallyUnimodular_iff, isTotallyUnimodular_iff] - constructor <;> intro hA k f g - · exact hA k (Sum.inl ∘ f) g - · if zerow : ∃ i, ∃ x', f i = Sum.inr x' then - obtain ⟨i, _, _⟩ := zerow - use 0 - rw [eq_comm] - apply det_eq_zero_of_row_eq_zero i - intro - simp_all - else - obtain ⟨_, rfl⟩ : ∃ f₀ : Fin k → m, f = Sum.inl ∘ f₀ := by - have hi (i : Fin k) : ∃ x, f i = Sum.inl x := - match hfi : f i with - | .inl x => ⟨x, rfl⟩ - | .inr x => (zerow ⟨i, x, hfi⟩).elim - choose f₀ hf₀ using hi - use f₀ - ext - apply hf₀ +lemma IsTotallyUnimodular.reindex {A : Matrix m n R} (em : m ≃ m') (en : n ≃ n') + (hA : A.IsTotallyUnimodular) : + (A.reindex em en).IsTotallyUnimodular := + hA.submatrix _ _ + +lemma reindex_isTotallyUnimodular (A : Matrix m n R) (em : m ≃ m') (en : n ≃ n') : + (A.reindex em en).IsTotallyUnimodular ↔ A.IsTotallyUnimodular := + ⟨fun hA => by simpa [Equiv.symm_apply_eq] using hA.reindex em.symm en.symm, + fun hA => hA.reindex _ _⟩ + +/-- If `A` is totally unimodular and each row of `B` is all zeros except for at most a single `1` or +a single `-1` then `fromRows A B` is totally unimodular. -/ +lemma IsTotallyUnimodular.fromRows_unitlike [DecidableEq n] {A : Matrix m n R} {B : Matrix m' n R} + (hA : A.IsTotallyUnimodular) + (hB : Nonempty n → ∀ i : m', ∃ j : n, ∃ s : SignType, B i = Pi.single j s.cast) : + (fromRows A B).IsTotallyUnimodular := by + intro k f g hf hg + induction k with + | zero => use 1; simp + | succ k ih => + specialize hB ⟨g 0⟩ + -- Either `f` is `inr` somewhere or `inl` everywhere + obtain ⟨i, j, hfi⟩ | ⟨f', rfl⟩ : (∃ i j, f i = .inr j) ∨ (∃ f', f = .inl ∘ f') := by + simp_rw [← Sum.isRight_iff, or_iff_not_imp_left, not_exists, Bool.not_eq_true, + Sum.isRight_eq_false, Sum.isLeft_iff] + intro hfr + choose f' hf' using hfr + exact ⟨f', funext hf'⟩ + · have hAB := det_succ_row ((fromRows A B).submatrix f g) i + simp only [submatrix_apply, hfi, fromRows_apply_inr] at hAB + obtain ⟨j', s, hj'⟩ := hB j + · simp only [hj', Function.update_apply] at hAB + by_cases hj'' : ∃ x, g x = j' + · obtain ⟨x, rfl⟩ := hj'' + rw [Fintype.sum_eq_single x fun y hxy => ?_, Pi.single_eq_same] at hAB + · rw [hAB] + change _ ∈ MonoidHom.mrange SignType.castHom.toMonoidHom + refine mul_mem (mul_mem ?_ (Set.mem_range_self s)) ?_ + · apply pow_mem + exact ⟨-1, by simp⟩ + · exact ih _ _ + (hf.comp Fin.succAbove_right_injective) + (hg.comp Fin.succAbove_right_injective) + · simp [Pi.single_eq_of_ne, hg.ne_iff.mpr hxy] + · rw [not_exists] at hj'' + use 0 + simpa [hj''] using hAB.symm + · rw [isTotallyUnimodular_iff] at hA apply hA -lemma fromColumns_col0_isTotallyUnimodular_iff (A : Matrix m n R) {n' : Type*} : +/-- If `A` is totally unimodular and each row of `B` is all zeros except for at most a single `1`, +then `fromRows A B` is totally unimodular. -/ +lemma fromRows_isTotallyUnimodular_iff_rows [DecidableEq n] {A : Matrix m n R} {B : Matrix m' n R} + (hB : Nonempty n → ∀ i : m', ∃ j : n, ∃ s : SignType, B i = Pi.single j s.cast) : + (fromRows A B).IsTotallyUnimodular ↔ A.IsTotallyUnimodular := + ⟨.submatrix Sum.inl id, fun hA => hA.fromRows_unitlike hB⟩ + +lemma fromRows_one_isTotallyUnimodular_iff [DecidableEq n] (A : Matrix m n R) : + (fromRows A (1 : Matrix n n R)).IsTotallyUnimodular ↔ A.IsTotallyUnimodular := + fromRows_isTotallyUnimodular_iff_rows <| fun h i ↦ + ⟨i, 1, funext fun j ↦ by simp [one_apply, Pi.single_apply, eq_comm]⟩ + +lemma one_fromRows_isTotallyUnimodular_iff [DecidableEq n] (A : Matrix m n R) : + (fromRows (1 : Matrix n n R) A).IsTotallyUnimodular ↔ A.IsTotallyUnimodular := by + have hA : + fromRows (1 : Matrix n n R) A = + (fromRows A (1 : Matrix n n R)).reindex (Equiv.sumComm m n) (Equiv.refl n) := by + aesop + rw [hA, reindex_isTotallyUnimodular, fromRows_one_isTotallyUnimodular_iff] + +lemma fromColumns_one_isTotallyUnimodular_iff [DecidableEq m] (A : Matrix m n R) : + (fromColumns A (1 : Matrix m m R)).IsTotallyUnimodular ↔ A.IsTotallyUnimodular := by + rw [←transpose_isTotallyUnimodular_iff, transpose_fromColumns, transpose_one, + fromRows_one_isTotallyUnimodular_iff, transpose_isTotallyUnimodular_iff] + +lemma one_fromColumns_isTotallyUnimodular_iff [DecidableEq m] (A : Matrix m n R) : + (fromColumns (1 : Matrix m m R) A).IsTotallyUnimodular ↔ A.IsTotallyUnimodular := by + rw [←transpose_isTotallyUnimodular_iff, transpose_fromColumns, transpose_one, + one_fromRows_isTotallyUnimodular_iff, transpose_isTotallyUnimodular_iff] + +alias ⟨_, IsTotallyUnimodular.fromRows_one⟩ := fromRows_one_isTotallyUnimodular_iff +alias ⟨_, IsTotallyUnimodular.one_fromRows⟩ := one_fromRows_isTotallyUnimodular_iff +alias ⟨_, IsTotallyUnimodular.fromColumns_one⟩ := fromColumns_one_isTotallyUnimodular_iff +alias ⟨_, IsTotallyUnimodular.one_fromColumns⟩ := one_fromColumns_isTotallyUnimodular_iff + +lemma fromRows_row0_isTotallyUnimodular_iff (A : Matrix m n R) : + (fromRows A (row m' 0)).IsTotallyUnimodular ↔ A.IsTotallyUnimodular := by + classical + refine fromRows_isTotallyUnimodular_iff_rows <| fun _ _ => ?_ + inhabit n + refine ⟨default, 0, ?_⟩ + ext x + simp [Pi.single_apply] + +lemma fromColumns_col0_isTotallyUnimodular_iff (A : Matrix m n R) : (fromColumns A (col n' 0)).IsTotallyUnimodular ↔ A.IsTotallyUnimodular := by rw [← transpose_isTotallyUnimodular_iff, transpose_fromColumns, transpose_col, fromRows_row0_isTotallyUnimodular_iff, transpose_isTotallyUnimodular_iff] From dc8397f7ed02f24627a498cea628c2784e7a1bbe Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 28 Nov 2024 13:32:40 +0000 Subject: [PATCH 368/829] feat: `(t \ s).ncard = t.ncard - s.ncard` in a ring (#19407) Version of `Finset.cast_card_sdiff` for `Set.ncard`. Also generalise `ncard_diff`. From Kneser (LeanCamCombi) --- Mathlib/Data/Set/Card.lean | 11 +++++++++-- 1 file changed, 9 insertions(+), 2 deletions(-) diff --git a/Mathlib/Data/Set/Card.lean b/Mathlib/Data/Set/Card.lean index e4c5eaa625091..e6992277d23f0 100644 --- a/Mathlib/Data/Set/Card.lean +++ b/Mathlib/Data/Set/Card.lean @@ -505,6 +505,7 @@ theorem encard_le_coe_iff_finite_ncard_le {k : ℕ} : s.encard ≤ k ↔ s.Finit theorem Infinite.ncard (hs : s.Infinite) : s.ncard = 0 := by rw [← Nat.card_coe_set_eq, @Nat.card_eq_zero_of_infinite _ hs.to_subtype] +@[gcongr] theorem ncard_le_ncard (hst : s ⊆ t) (ht : t.Finite := by toFinite_tac) : s.ncard ≤ t.ncard := by rw [← Nat.cast_le (α := ℕ∞), ht.cast_ncard_eq, (ht.subset hst).cast_ncard_eq] @@ -827,9 +828,15 @@ theorem ncard_diff_add_ncard_of_subset (h : s ⊆ t) (ht : t.Finite := by toFini rw [ht.cast_ncard_eq, (ht.subset h).cast_ncard_eq, (ht.diff _).cast_ncard_eq, encard_diff_add_encard_of_subset h] -theorem ncard_diff (h : s ⊆ t) (ht : t.Finite := by toFinite_tac) : +theorem ncard_diff (hst : s ⊆ t) (hs : s.Finite := by toFinite_tac) : (t \ s).ncard = t.ncard - s.ncard := by - rw [← ncard_diff_add_ncard_of_subset h ht, add_tsub_cancel_right] + obtain ht | ht := t.finite_or_infinite + · rw [← ncard_diff_add_ncard_of_subset hst ht, add_tsub_cancel_right] + · rw [ht.ncard, Nat.zero_sub, (ht.diff hs).ncard] + +lemma cast_ncard_sdiff {R : Type*} [AddGroupWithOne R] (hst : s ⊆ t) (ht : t.Finite) : + ((t \ s).ncard : R) = t.ncard - s.ncard := by + rw [ncard_diff hst (ht.subset hst), Nat.cast_sub (ncard_le_ncard hst ht)] theorem ncard_le_ncard_diff_add_ncard (s t : Set α) (ht : t.Finite := by toFinite_tac) : s.ncard ≤ (s \ t).ncard + t.ncard := by From 44a42c7c27742d343ffe2e9b66e47f935901d236 Mon Sep 17 00:00:00 2001 From: Xavier Roblot Date: Thu, 28 Nov 2024 13:58:22 +0000 Subject: [PATCH 369/829] feat(Analysis/BoxIntegral): Add BoxPartition defined by subdivision of the integer lattice (#16675) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Let `n` be a positive integer. This PR defines `BoxIntegral.unitPartition.Box` which are boxes in `ι → ℝ` obtained by dividing evenly the boxes of the integer lattice`ι → ℤ` into smaller boxes of sides of length `1 / n`. These boxes will be used in another PR #12405 that relates point counting in sets and integration. This PR is part of the proof of the Analytic Class Number Formula. Co-authored-by: Xavier Roblot <46200072+xroblot@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/Analysis/BoxIntegral/Box/Basic.lean | 7 + .../Analysis/BoxIntegral/UnitPartition.lean | 296 ++++++++++++++++++ 3 files changed, 304 insertions(+) create mode 100644 Mathlib/Analysis/BoxIntegral/UnitPartition.lean diff --git a/Mathlib.lean b/Mathlib.lean index 2ac571937ef73..b75c82cf7623e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1086,6 +1086,7 @@ import Mathlib.Analysis.BoxIntegral.Partition.Measure import Mathlib.Analysis.BoxIntegral.Partition.Split import Mathlib.Analysis.BoxIntegral.Partition.SubboxInduction import Mathlib.Analysis.BoxIntegral.Partition.Tagged +import Mathlib.Analysis.BoxIntegral.UnitPartition import Mathlib.Analysis.CStarAlgebra.ApproximateUnit import Mathlib.Analysis.CStarAlgebra.Basic import Mathlib.Analysis.CStarAlgebra.Classes diff --git a/Mathlib/Analysis/BoxIntegral/Box/Basic.lean b/Mathlib/Analysis/BoxIntegral/Box/Basic.lean index 230e52b1389bb..fa0d86469c73b 100644 --- a/Mathlib/Analysis/BoxIntegral/Box/Basic.lean +++ b/Mathlib/Analysis/BoxIntegral/Box/Basic.lean @@ -208,6 +208,13 @@ theorem monotone_upper : Monotone fun I : Box ι ↦ I.upper := theorem coe_subset_Icc : ↑I ⊆ Box.Icc I := fun _ hx ↦ ⟨fun i ↦ (hx i).1.le, fun i ↦ (hx i).2⟩ +theorem isBounded_Icc [Finite ι] (I : Box ι) : Bornology.IsBounded (Box.Icc I) := by + cases nonempty_fintype ι + exact Metric.isBounded_Icc _ _ + +theorem isBounded [Finite ι] (I : Box ι) : Bornology.IsBounded I.toSet := + Bornology.IsBounded.subset I.isBounded_Icc coe_subset_Icc + /-! ### Supremum of two boxes -/ diff --git a/Mathlib/Analysis/BoxIntegral/UnitPartition.lean b/Mathlib/Analysis/BoxIntegral/UnitPartition.lean new file mode 100644 index 0000000000000..092ae83e8c61f --- /dev/null +++ b/Mathlib/Analysis/BoxIntegral/UnitPartition.lean @@ -0,0 +1,296 @@ +/- +Copyright (c) 2024 Xavier Roblot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Xavier Roblot +-/ +import Mathlib.Analysis.BoxIntegral.Partition.Measure +import Mathlib.Analysis.BoxIntegral.Partition.Tagged + +/-! +# Unit Partition + +Fix `n` a positive integer. `BoxIntegral.unitPartition.box` are boxes in `ι → ℝ` obtained by +dividing the unit box uniformly into boxes of side length `1 / n` and translating the boxes by +the lattice `ι → ℤ` so that they cover the whole space. For fixed `n`, there are indexed by +vectors `ν : ι → ℤ`. + +Let `B` be a `BoxIntegral`. A `unitPartition.box` is admissible for `B` (more precisely its index is +admissible) if it is contained in `B`. There are finitely many admissible `unitPartition.box` for +`B` and thus we can form the corresponing tagged prepartition, see +`BoxIntegral.unitPartition.prepartition` (note that each `unitPartition.Box` comes with its +tag situated at its "upper most" vertex). If `B` satisfies `hasIntegralVertices`, that +is its vertices are in `ι → ℤ`, then the corresponding prepartition is actually a partition. + +## Main definitions and results + +* `BoxIntegral.hasIntegralVertices`: a `Prop` that states that the vertices of the box have +coordinates in `ℤ` + +* `BoxIntegral.unitPartition.box`: a `BoxIntegral`, indexed by `ν : ι → ℤ`, with vertices +`ν i / n` and of side length `1 / n`. + +* `BoxIntegral.unitPartition.admissibleIndex`: For `B : BoxIntegral.Box`, the set of indices of +`unitPartition.box` that are subsets of `B`. This is a finite set. + +* `BoxIntegral.unitPartition.prepartition_isPartition`: For `B : BoxIntegral.Box`, if `B` +has integral vertices, then the prepartition of `unitPartition.box` admissible for `B` is a +partition of `B`. +-/ + +noncomputable section + +variable {ι : Type*} + +section hasIntegralVertices + +open Bornology + +/-- A `BoxIntegral.Box` has integral vertices if its vertices have coordinates in `ℤ`. -/ +def BoxIntegral.hasIntegralVertices (B : Box ι) : Prop := + ∃ l u : ι → ℤ, (∀ i, B.lower i = l i) ∧ (∀ i, B.upper i = u i) + +/-- Any bounded set is contained in a `BoxIntegral.Box` with integral vertices. -/ +theorem BoxIntegral.le_hasIntegralVertices_of_isBounded [Finite ι] {s : Set (ι → ℝ)} + (h : IsBounded s) : + ∃ B : BoxIntegral.Box ι, hasIntegralVertices B ∧ s ≤ B := by + have := Fintype.ofFinite ι + obtain ⟨R, hR₁, hR₂⟩ := IsBounded.subset_ball_lt h 0 0 + let C : ℕ := ⌈R⌉₊ + have hC := Nat.ceil_pos.mpr hR₁ + let I : Box ι := Box.mk (fun _ ↦ - C) (fun _ ↦ C ) + (fun _ ↦ by simp only [neg_lt_self_iff, Nat.cast_pos, hC]) + refine ⟨I, ⟨fun _ ↦ - C, fun _ ↦ C, fun i ↦ (Int.cast_neg_natCast C).symm, fun _ ↦ rfl⟩, + le_trans hR₂ ?_⟩ + suffices Metric.ball (0 : ι → ℝ) C ≤ I from + le_trans (Metric.ball_subset_ball (Nat.le_ceil R)) this + intro x hx + simp_rw [mem_ball_zero_iff, pi_norm_lt_iff (Nat.cast_pos.mpr hC), Real.norm_eq_abs, abs_lt] at hx + exact fun i ↦ ⟨(hx i).1, le_of_lt (hx i).2⟩ + +end hasIntegralVertices + +namespace BoxIntegral.unitPartition + +open Bornology MeasureTheory Fintype BoxIntegral + +variable (n : ℕ) + +/-- A `BoxIntegral`, indexed by a positive integer `n` and `ν : ι → ℤ`, with corners `ν i / n` +and of side length `1 / n`. -/ +def box [NeZero n] (ν : ι → ℤ) : Box ι where + lower := fun i ↦ ν i / n + upper := fun i ↦ (ν i + 1) / n + lower_lt_upper := fun _ ↦ by norm_num [add_div, n.pos_of_neZero] + +@[simp] +theorem box_lower [NeZero n] (ν : ι → ℤ) : + (box n ν).lower = fun i ↦ (ν i / n : ℝ) := rfl + +@[simp] +theorem box_upper [NeZero n] (ν : ι → ℤ) : + (box n ν).upper = fun i ↦ ((ν i + 1)/ n : ℝ) := rfl + +variable {n} in +@[simp] +theorem mem_box_iff [NeZero n] {ν : ι → ℤ} {x : ι → ℝ} : + x ∈ box n ν ↔ ∀ i, ν i / n < x i ∧ x i ≤ (ν i + 1) / n := by + simp_rw [Box.mem_def, box, Set.mem_Ioc] + +variable {n} in +theorem mem_box_iff' [NeZero n] {ν : ι → ℤ} {x : ι → ℝ} : + x ∈ box n ν ↔ ∀ i, ν i < n * x i ∧ n * x i ≤ ν i + 1 := by + have h : 0 < (n : ℝ) := Nat.cast_pos.mpr <| n.pos_of_neZero + simp_rw [mem_box_iff, ← _root_.le_div_iff₀' h, ← div_lt_iff₀' h] + +/-- The tag of (the index of) a `unitPartition.box`. -/ +abbrev tag (ν : ι → ℤ) : ι → ℝ := fun i ↦ (ν i + 1) / n + +@[simp] +theorem tag_apply (ν : ι → ℤ) (i : ι) : tag n ν i = (ν i + 1) / n := rfl + +variable [NeZero n] + +theorem tag_injective : Function.Injective (fun ν : ι → ℤ ↦ tag n ν) := by + refine fun _ _ h ↦ funext_iff.mpr fun i ↦ ?_ + have := congr_arg (fun x ↦ x i) h + simp_rw [tag_apply, div_left_inj' (c := (n : ℝ)) (Nat.cast_ne_zero.mpr (NeZero.ne n)), + add_left_inj, Int.cast_inj] at this + exact this + +theorem tag_mem (ν : ι → ℤ) : + tag n ν ∈ box n ν := by + refine mem_box_iff.mpr fun _ ↦ ?_ + rw [tag, add_div] + have h : 0 < (n : ℝ) := Nat.cast_pos.mpr <| n.pos_of_neZero + exact ⟨lt_add_of_pos_right _ (by positivity), le_rfl⟩ + +/-- For `x : ι → ℝ`, its index is the index of the unique `unitPartition.box` to which +it belongs. -/ +def index (x : ι → ℝ) (i : ι) : ℤ := ⌈n * x i⌉ - 1 + +@[simp] +theorem index_apply (m : ℕ) {x : ι → ℝ} (i : ι) : + index m x i = ⌈m * x i⌉ - 1 := rfl + +variable {n} in +theorem mem_box_iff_index {x : ι → ℝ} {ν : ι → ℤ} : + x ∈ box n ν ↔ index n x = ν := by + simp_rw [mem_box_iff', funext_iff, index_apply, sub_eq_iff_eq_add, Int.ceil_eq_iff, + Int.cast_add, Int.cast_one, add_sub_cancel_right] + +@[simp] +theorem index_tag (ν : ι → ℤ) : + index n (tag n ν) = ν := mem_box_iff_index.mp (tag_mem n ν) + +variable {n} in +theorem disjoint {ν ν' : ι → ℤ} : + ν ≠ ν' ↔ Disjoint (box n ν).toSet (box n ν').toSet := by + rw [not_iff_comm, Set.not_disjoint_iff] + refine ⟨fun ⟨x, hx, hx'⟩ ↦ ?_, fun h ↦ ⟨tag n ν, tag_mem n ν, h ▸ tag_mem n ν⟩⟩ + rw [← mem_box_iff_index.mp hx, ← mem_box_iff_index.mp hx'] + +theorem box_injective : Function.Injective (fun ν : ι → ℤ ↦ box n ν) := by + intro _ _ h + contrapose! h + exact Box.ne_of_disjoint_coe (disjoint.mp h) + +lemma box.upper_sub_lower (ν : ι → ℤ) (i : ι) : + (box n ν ).upper i - (box n ν).lower i = 1 / n := by + simp_rw [box, add_div, add_sub_cancel_left] + +variable [Fintype ι] + +theorem diam_boxIcc (ν : ι → ℤ) : + Metric.diam (Box.Icc (box n ν)) ≤ 1 / n := by + rw [BoxIntegral.Box.Icc_eq_pi] + refine ENNReal.toReal_le_of_le_ofReal (by positivity) <| EMetric.diam_pi_le_of_le (fun i ↦ ?_) + simp_rw [Real.ediam_Icc, box.upper_sub_lower, le_rfl] + +@[simp] +theorem volume_box (ν : ι → ℤ) : + volume (box n ν : Set (ι → ℝ)) = 1 / n ^ card ι := by + simp_rw [volume_pi, BoxIntegral.Box.coe_eq_pi, Measure.pi_pi, Real.volume_Ioc, + box.upper_sub_lower, Finset.prod_const, ENNReal.ofReal_div_of_pos (Nat.cast_pos.mpr + n.pos_of_neZero), ENNReal.ofReal_one, ENNReal.ofReal_natCast, one_div, ENNReal.inv_pow, + Finset.card_univ] + +theorem setFinite_index {s : Set (ι → ℝ)} (hs₁ : NullMeasurableSet s) (hs₂ : volume s ≠ ⊤) : + Set.Finite {ν : ι → ℤ | ↑(box n ν) ⊆ s} := by + refine (Measure.finite_const_le_meas_of_disjoint_iUnion₀ volume (ε := 1 / n ^ card ι) + (by norm_num) (As := fun ν : ι → ℤ ↦ (box n ν) ∩ s) (fun ν ↦ ?_) (fun _ _ h ↦ ?_) ?_).subset + (fun _ hν ↦ ?_) + · refine NullMeasurableSet.inter ?_ hs₁ + exact (box n ν).measurableSet_coe.nullMeasurableSet + · exact ((Disjoint.inter_right _ (disjoint.mp h)).inter_left _ ).aedisjoint + · exact lt_top_iff_ne_top.mp <| measure_lt_top_of_subset + (by simp only [Set.iUnion_subset_iff, Set.inter_subset_right, implies_true]) hs₂ + · rw [Set.mem_setOf, Set.inter_eq_self_of_subset_left hν, volume_box] + +/-- For `B : BoxIntegral.Box`, the set of indices of `unitPartition.Box` that are subsets of `B`. +This is a finite set. These boxes cover `B` if it has integral vertices, see +`unitPartition.prepartition_isPartition`. -/ +def admissibleIndex (B : Box ι) : Finset (ι → ℤ) := by + refine (setFinite_index n B.measurableSet_coe.nullMeasurableSet ?_).toFinset + exact lt_top_iff_ne_top.mp (IsBounded.measure_lt_top B.isBounded) + +variable {n} in +theorem mem_admissibleIndex_iff {B : Box ι} {ν : ι → ℤ} : + ν ∈ admissibleIndex n B ↔ box n ν ≤ B := by + rw [admissibleIndex, Set.Finite.mem_toFinset, Set.mem_setOf_eq, Box.coe_subset_coe] + +open Classical in +/-- For `B : BoxIntegral.Box`, the `TaggedPrepartition` formed by the set of all +`unitPartition.box` whose index is `B`-admissible. -/ +def prepartition (B : Box ι) : TaggedPrepartition B where + boxes := Finset.image (fun ν ↦ box n ν) (admissibleIndex n B) + le_of_mem' _ hI := by + obtain ⟨_, hν, rfl⟩ := Finset.mem_image.mp hI + exact mem_admissibleIndex_iff.mp hν + pairwiseDisjoint _ hI₁ _ hI₂ h := by + obtain ⟨_, _, rfl⟩ := Finset.mem_image.mp hI₁ + obtain ⟨_, _, rfl⟩ := Finset.mem_image.mp hI₂ + exact disjoint.mp fun x ↦ h (congrArg (box n) x) + tag I := + if hI : ∃ ν ∈ admissibleIndex n B, I = box n ν then tag n hI.choose else B.exists_mem.choose + tag_mem_Icc I := by + by_cases hI : ∃ ν ∈ admissibleIndex n B, I = box n ν + · simp_rw [dif_pos hI] + exact Box.coe_subset_Icc <| (mem_admissibleIndex_iff.mp hI.choose_spec.1) (tag_mem n _) + · simp_rw [dif_neg hI] + exact Box.coe_subset_Icc B.exists_mem.choose_spec + +variable {n} in +@[simp] +theorem mem_prepartition_iff {B I : Box ι} : + I ∈ prepartition n B ↔ ∃ ν ∈ admissibleIndex n B, box n ν = I := by + classical + rw [prepartition, TaggedPrepartition.mem_mk, Prepartition.mem_mk, Finset.mem_image] + +variable {n} in +theorem mem_prepartition_boxes_iff {B I : Box ι} : + I ∈ (prepartition n B).boxes ↔ ∃ ν ∈ admissibleIndex n B, box n ν = I := + mem_prepartition_iff + +theorem prepartition_tag {ν : ι → ℤ} {B : Box ι} (hν : ν ∈ admissibleIndex n B) : + (prepartition n B).tag (box n ν) = tag n ν := by + dsimp only [prepartition] + have h : ∃ ν' ∈ admissibleIndex n B, box n ν = box n ν' := ⟨ν, hν, rfl⟩ + rw [dif_pos h, (tag_injective n).eq_iff, ← (box_injective n).eq_iff] + exact h.choose_spec.2.symm + +theorem box_index_tag_eq_self {B I : Box ι} (hI : I ∈ (prepartition n B).boxes) : + box n (index n ((prepartition n B).tag I)) = I := by + obtain ⟨ν, hν, rfl⟩ := mem_prepartition_boxes_iff.mp hI + rw [prepartition_tag n hν, index_tag] + +theorem prepartition_isHenstock (B : Box ι) : + (prepartition n B).IsHenstock := by + intro _ hI + obtain ⟨ν, hν, rfl⟩ := mem_prepartition_iff.mp hI + rw [prepartition_tag n hν] + exact Box.coe_subset_Icc (tag_mem _ _) + +theorem prepartition_isSubordinate (B : Box ι) {r : ℝ} (hr : 0 < r) (hn : 1 / n ≤ r) : + (prepartition n B).IsSubordinate (fun _ ↦ ⟨r, hr⟩) := by + intro _ hI + obtain ⟨ν, hν, rfl⟩ := mem_prepartition_iff.mp hI + refine fun _ h ↦ le_trans (Metric.dist_le_diam_of_mem (Box.isBounded_Icc _) h ?_) ?_ + · rw [prepartition_tag n hν] + exact Box.coe_subset_Icc (tag_mem _ _) + · exact le_trans (diam_boxIcc n ν) hn + +private theorem mem_admissibleIndex_of_mem_box_aux₁ (x : ℝ) (a : ℤ) : + a < x ↔ a ≤ (⌈n * x⌉ - 1) / (n : ℝ) := by + have h : 0 < (n : ℝ) := Nat.cast_pos.mpr <| n.pos_of_neZero + rw [le_div_iff₀' h, le_sub_iff_add_le, + show (n : ℝ) * a + 1 = (n * a + 1 : ℤ) by norm_cast, + Int.cast_le, Int.add_one_le_ceil_iff, Int.cast_mul, Int.cast_natCast, mul_lt_mul_left h] + +private theorem mem_admissibleIndex_of_mem_box_aux₂ (x : ℝ) (a : ℤ) : + x ≤ a ↔ (⌈n * x⌉ - 1 + 1) / (n : ℝ) ≤ a := by + have h : 0 < (n : ℝ) := Nat.cast_pos.mpr <| n.pos_of_neZero + rw [sub_add_cancel, div_le_iff₀' h, + show (n : ℝ) * a = (n * a : ℤ) by norm_cast, + Int.cast_le, Int.ceil_le, Int.cast_mul, Int.cast_natCast, mul_le_mul_left h] + +/-- If `B : BoxIntegral.Box` has integral vertices and contains the point `x`, then the index of +`x` is admissible for `B`. -/ +theorem mem_admissibleIndex_of_mem_box {B : Box ι} (hB : hasIntegralVertices B) {x : ι → ℝ} + (hx : x ∈ B) : index n x ∈ admissibleIndex n B := by + obtain ⟨l, u, hl, hu⟩ := hB + simp_rw [mem_admissibleIndex_iff, Box.le_iff_bounds, box_lower, box_upper, Pi.le_def, + index_apply, hl, hu, ← forall_and] + push_cast + refine fun i ↦ ⟨?_, ?_⟩ + · exact (mem_admissibleIndex_of_mem_box_aux₁ n (x i) (l i)).mp ((hl i) ▸ (hx i).1) + · exact (mem_admissibleIndex_of_mem_box_aux₂ n (x i) (u i)).mp ((hu i) ▸ (hx i).2) + +/-- If `B : BoxIntegral.Box` has integral vertices, then `prepartition n B` is a partition of +`B`. -/ +theorem prepartition_isPartition {B : Box ι} (hB : hasIntegralVertices B) : + (prepartition n B).IsPartition := by + refine fun x hx ↦ ⟨box n (index n x), ?_, mem_box_iff_index.mpr rfl⟩ + rw [TaggedPrepartition.mem_toPrepartition, mem_prepartition_iff] + exact ⟨index n x, mem_admissibleIndex_of_mem_box n hB hx, rfl⟩ + +end BoxIntegral.unitPartition From 9d44280a87c6be2122cdeb5b95fe0c18306b050e Mon Sep 17 00:00:00 2001 From: David Renshaw Date: Thu, 28 Nov 2024 13:58:24 +0000 Subject: [PATCH 370/829] doc: update README.md instructions for building documentation locally (#18713) Since #14229, the current instructions in README.md for building docs no longer work. This PR updates the instructions to refer to the `mathlib4_docs` repo, which is discussed in [this zulip thread](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/How.20does.20build.20.26.20deploy.20document.20for.20Mathlib4.20currently.3F/near/476591680). --- README.md | 17 ++++++++++------- 1 file changed, 10 insertions(+), 7 deletions(-) diff --git a/README.md b/README.md index b93dbbdcabe09..6e5a49dbeafd5 100644 --- a/README.md +++ b/README.md @@ -88,17 +88,20 @@ Call `lake exe cache` to see its help menu. ### Building HTML documentation -Building HTML documentation locally is straightforward, but it may take a while (>20 minutes): +The [mathlib4_docs repository](https://github.com/leanprover-community/mathlib4_docs) +is responsible for generating and publishing the +[mathlib4 docs](https://leanprover-community.github.io/mathlib4_docs/index.html). +That repo can be used to build the docs locally: ```shell -lake -R -Kdoc=on update doc-gen4 +git clone https://github.com/leanprover-community/mathlib4_docs.git +cd mathlib4_docs +cp ../mathlib4/lean-toolchain . +lake exe cache get lake build Mathlib:docs ``` - -The HTML files can then be found in `build/doc`. - -Warning: these commands will make a change to `lake-manifest.json` -which should *not* be committed to Mathlib. +The last step may take a while (>20 minutes). +The HTML files can then be found in `.lake/build/doc`. ## Transitioning from Lean 3 From 63a0f6a0a94a4f16eff2c040076138cac61c58d2 Mon Sep 17 00:00:00 2001 From: David Renshaw Date: Thu, 28 Nov 2024 13:58:25 +0000 Subject: [PATCH 371/829] feat: add Nat.digits_base_pow_mul and Nat.digits_append_zeroes_append_digits (#19179) Adds two `Nat.digits` lemmas that I needed in Compfiles for USAMO 1992 Problem 1: https://github.com/dwrensha/compfiles/blob/472f49e2b567e05c2cc1d963d5efbd6118f2b66c/Compfiles/Usa1992P1.lean#L32-L52 `Nat.digits_base_pow_mul` is like [`Nat.digits_append_digits`](https://github.com/leanprover-community/mathlib4/blob/7aeae34dfe1ca991bdb678b042ab1c7e3bea0df7/Mathlib/Data/Nat/Digits.lean#L413-L414), but with `Nat.digits b n` replaced by `List.replicate k 0`. This new lemma is useful because `Nat.digits b n` can never equal a nonempty list of all zeroes. `Nat.digits_append_zeroes_append_digits` is like `Nat.digits_append_digits` with an extra `List.replicate k 0` inserted in the middle. --- Mathlib/Data/Nat/Digits.lean | 20 ++++++++++++++++++++ 1 file changed, 20 insertions(+) diff --git a/Mathlib/Data/Nat/Digits.lean b/Mathlib/Data/Nat/Digits.lean index 4ba490f6c3ada..f9385aca645c1 100644 --- a/Mathlib/Data/Nat/Digits.lean +++ b/Mathlib/Data/Nat/Digits.lean @@ -406,6 +406,19 @@ theorem lt_base_pow_length_digits {b m : ℕ} (hb : 1 < b) : m < b ^ (digits b m rcases b with (_ | _ | b) <;> try simp_all exact lt_base_pow_length_digits' +theorem digits_base_pow_mul {b k m : ℕ} (hb : 1 < b) (hm : 0 < m) : + digits b (b ^ k * m) = List.replicate k 0 ++ digits b m := by + induction k generalizing m with + | zero => simp + | succ k ih => + have hmb : 0 < m * b := lt_mul_of_lt_of_one_lt' hm hb + let h1 := digits_def' hb hmb + have h2 : m = m * b / b := + Nat.eq_div_of_mul_eq_left (not_eq_zero_of_lt hb) rfl + simp only [mul_mod_left, ← h2] at h1 + rw [List.replicate_succ', List.append_assoc, List.singleton_append, ← h1, ← ih hmb] + ring_nf + theorem ofDigits_digits_append_digits {b m n : ℕ} : ofDigits b (digits b n ++ digits b m) = n + b ^ (digits b n).length * m := by rw [ofDigits_append, ofDigits_digits, ofDigits_digits] @@ -423,6 +436,13 @@ theorem digits_append_digits {b m n : ℕ} (hb : 0 < b) : · exact (List.getLast_append' _ _ h) ▸ (getLast_digit_ne_zero _ <| digits_ne_nil_iff_ne_zero.mp h) +theorem digits_append_zeroes_append_digits {b k m n : ℕ} (hb : 1 < b) (hm : 0 < m) : + digits b n ++ List.replicate k 0 ++ digits b m = + digits b (n + b ^ ((digits b n).length + k) * m) := by + rw [List.append_assoc, ← digits_base_pow_mul hb hm] + simp only [digits_append_digits (zero_lt_of_lt hb), digits_inj_iff, add_right_inj] + ring + theorem digits_len_le_digits_len_succ (b n : ℕ) : (digits b n).length ≤ (digits b (n + 1)).length := by rcases Decidable.eq_or_ne n 0 with (rfl | hn) From d6b454822f968c3dc4f159167e124b0463e8c6f6 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Thu, 28 Nov 2024 13:58:27 +0000 Subject: [PATCH 372/829] chore: split out Algebra.GroupWithZero.Nat (#19342) This means `Data.Nat.Prime.Defs` no longer depends on rings. --- Mathlib.lean | 1 + Mathlib/Algebra/EuclideanDomain/Int.lean | 1 - Mathlib/Algebra/Group/Nat/Basic.lean | 4 ++ Mathlib/Algebra/GroupWithZero/Nat.lean | 53 +++++++++++++++++++++++ Mathlib/Algebra/Ring/Nat.lean | 54 +++++++++++++----------- Mathlib/Data/Nat/GCD/Basic.lean | 4 +- Mathlib/Data/Nat/ModEq.lean | 1 + Mathlib/Data/Nat/Prime/Defs.lean | 8 ++-- MathlibTest/linear_combination.lean | 2 +- MathlibTest/rewrites.lean | 1 + 10 files changed, 96 insertions(+), 33 deletions(-) create mode 100644 Mathlib/Algebra/GroupWithZero/Nat.lean diff --git a/Mathlib.lean b/Mathlib.lean index b75c82cf7623e..742aee5ebc6c8 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -365,6 +365,7 @@ import Mathlib.Algebra.GroupWithZero.Hom import Mathlib.Algebra.GroupWithZero.Indicator import Mathlib.Algebra.GroupWithZero.InjSurj import Mathlib.Algebra.GroupWithZero.Invertible +import Mathlib.Algebra.GroupWithZero.Nat import Mathlib.Algebra.GroupWithZero.NeZero import Mathlib.Algebra.GroupWithZero.NonZeroDivisors import Mathlib.Algebra.GroupWithZero.Opposite diff --git a/Mathlib/Algebra/EuclideanDomain/Int.lean b/Mathlib/Algebra/EuclideanDomain/Int.lean index 814a60c3f8170..cf153272fea14 100644 --- a/Mathlib/Algebra/EuclideanDomain/Int.lean +++ b/Mathlib/Algebra/EuclideanDomain/Int.lean @@ -6,7 +6,6 @@ Authors: Louis Carlin, Mario Carneiro import Mathlib.Algebra.EuclideanDomain.Defs import Mathlib.Algebra.Order.Group.Unbundled.Int import Mathlib.Algebra.Ring.Int.Defs -import Mathlib.Algebra.Ring.Nat /-! # Instances for Euclidean domains diff --git a/Mathlib/Algebra/Group/Nat/Basic.lean b/Mathlib/Algebra/Group/Nat/Basic.lean index 75212737a122b..9d677accc18c3 100644 --- a/Mathlib/Algebra/Group/Nat/Basic.lean +++ b/Mathlib/Algebra/Group/Nat/Basic.lean @@ -20,6 +20,10 @@ namespace Nat /-! ### Instances -/ +instance instMulOneClass : MulOneClass ℕ where + one_mul := Nat.one_mul + mul_one := Nat.mul_one + instance instAddCancelCommMonoid : AddCancelCommMonoid ℕ where add := Nat.add add_assoc := Nat.add_assoc diff --git a/Mathlib/Algebra/GroupWithZero/Nat.lean b/Mathlib/Algebra/GroupWithZero/Nat.lean new file mode 100644 index 0000000000000..da61db4e4caf8 --- /dev/null +++ b/Mathlib/Algebra/GroupWithZero/Nat.lean @@ -0,0 +1,53 @@ +/- +Copyright (c) 2014 Floris van Doorn (c) 2016 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro +-/ +import Mathlib.Algebra.Group.Nat.Basic +import Mathlib.Algebra.GroupWithZero.Defs +import Mathlib.Tactic.Spread + +/-! +# The natural numbers form a `CancelCommMonoidWithZero` + +This file contains the `CancelCommMonoidWithZero` instance on the natural numbers. + +See note [foundational algebra order theory]. +-/ + +assert_not_exists Ring + +namespace Nat + +instance instMulZeroClass : MulZeroClass ℕ where + zero_mul := Nat.zero_mul + mul_zero := Nat.mul_zero + +instance instSemigroupWithZero : SemigroupWithZero ℕ where + __ := instSemigroup + __ := instMulZeroClass + +instance instMonoidWithZero : MonoidWithZero ℕ where + __ := instMonoid + __ := instMulZeroClass + __ := instSemigroupWithZero + +instance instCommMonoidWithZero : CommMonoidWithZero ℕ where + __ := instCommMonoid + __ := instMonoidWithZero + +instance instIsLeftCancelMulZero : IsLeftCancelMulZero ℕ where + mul_left_cancel_of_ne_zero h := Nat.eq_of_mul_eq_mul_left (Nat.pos_of_ne_zero h) + +instance instCancelCommMonoidWithZero : CancelCommMonoidWithZero ℕ where + __ := instCommMonoidWithZero + __ := instIsLeftCancelMulZero + +instance instMulDivCancelClass : MulDivCancelClass ℕ where + mul_div_cancel _ _b hb := Nat.mul_div_cancel _ (Nat.pos_iff_ne_zero.2 hb) + +instance instMulZeroOneClass : MulZeroOneClass ℕ where + __ := instMulZeroClass + __ := instMulOneClass + +end Nat diff --git a/Mathlib/Algebra/Ring/Nat.lean b/Mathlib/Algebra/Ring/Nat.lean index e2c57fc380685..74467ee1cc0b8 100644 --- a/Mathlib/Algebra/Ring/Nat.lean +++ b/Mathlib/Algebra/Ring/Nat.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Mathlib.Algebra.CharZero.Defs -import Mathlib.Algebra.Group.Nat.Basic +import Mathlib.Algebra.GroupWithZero.Nat import Mathlib.Algebra.Ring.Defs /-! @@ -17,38 +17,42 @@ See note [foundational algebra order theory]. namespace Nat -/-! ### Instances -/ - -instance instCommSemiring : CommSemiring ℕ where - __ := instCommMonoid - left_distrib := Nat.left_distrib - right_distrib := Nat.right_distrib - zero_mul := Nat.zero_mul - mul_zero := Nat.mul_zero - npow m n := n ^ m - npow_zero := Nat.pow_zero - npow_succ _ _ := rfl +instance instAddMonoidWithOne : AddMonoidWithOne ℕ where natCast n := n natCast_zero := rfl natCast_succ _ := rfl -instance instCancelCommMonoidWithZero : CancelCommMonoidWithZero ℕ where - __ : CommMonoidWithZero ℕ := by infer_instance - mul_left_cancel_of_ne_zero h := Nat.eq_of_mul_eq_mul_left (Nat.pos_of_ne_zero h) +instance instAddCommMonoidWithOne : AddCommMonoidWithOne ℕ where + __ := instAddMonoidWithOne + __ := instAddCommMonoid -instance instMulDivCancelClass : MulDivCancelClass ℕ where - mul_div_cancel _ _b hb := Nat.mul_div_cancel _ (Nat.pos_iff_ne_zero.2 hb) +instance instDistrib : Distrib ℕ where + left_distrib := Nat.left_distrib + right_distrib := Nat.right_distrib -instance instCharZero : CharZero ℕ where cast_injective := Function.injective_id +instance instNonUnitalNonAssocSemiring : NonUnitalNonAssocSemiring ℕ where + __ := instAddCommMonoid + __ := instDistrib + __ := instMulZeroClass -/-! -### Extra instances to short-circuit type class resolution +instance instNonUnitalSemiring : NonUnitalSemiring ℕ where + __ := instNonUnitalNonAssocSemiring + __ := instSemigroupWithZero -These also prevent non-computable instances being used to construct these instances non-computably. --/ +instance instNonAssocSemiring : NonAssocSemiring ℕ where + __ := instNonUnitalNonAssocSemiring + __ := instMulZeroOneClass + __ := instAddCommMonoidWithOne -instance instAddCommMonoidWithOne : AddCommMonoidWithOne ℕ := by infer_instance -instance instDistrib : Distrib ℕ := by infer_instance -instance instSemiring : Semiring ℕ := by infer_instance +instance instSemiring : Semiring ℕ where + __ := instNonUnitalSemiring + __ := instNonAssocSemiring + __ := instMonoidWithZero + +instance instCommSemiring : CommSemiring ℕ where + __ := instSemiring + __ := instCommMonoid + +instance instCharZero : CharZero ℕ where cast_injective := Function.injective_id end Nat diff --git a/Mathlib/Data/Nat/GCD/Basic.lean b/Mathlib/Data/Nat/GCD/Basic.lean index 2adab886eac34..78cfdec77c654 100644 --- a/Mathlib/Data/Nat/GCD/Basic.lean +++ b/Mathlib/Data/Nat/GCD/Basic.lean @@ -6,7 +6,7 @@ Authors: Jeremy Avigad, Leonardo de Moura import Batteries.Data.Nat.Gcd import Mathlib.Algebra.Group.Nat.Units import Mathlib.Algebra.GroupWithZero.Divisibility -import Mathlib.Algebra.Ring.Nat +import Mathlib.Algebra.GroupWithZero.Nat /-! # Properties of `Nat.gcd`, `Nat.lcm`, and `Nat.Coprime` @@ -300,7 +300,7 @@ theorem Coprime.mul_add_mul_ne_mul {m n a b : ℕ} (cop : Coprime m n) (ha : a ((congr_arg _ h).mpr (Nat.dvd_mul_right m n))) rw [mul_comm, mul_ne_zero_iff, ← one_le_iff_ne_zero] at ha hb refine mul_ne_zero hb.2 ha.2 (eq_zero_of_mul_eq_self_left (ne_of_gt (add_le_add ha.1 hb.1)) ?_) - rw [← mul_assoc, ← h, add_mul, add_mul, mul_comm _ n, ← mul_assoc, mul_comm y] + rw [← mul_assoc, ← h, Nat.add_mul, Nat.add_mul, mul_comm _ n, ← mul_assoc, mul_comm y] variable {x n m : ℕ} diff --git a/Mathlib/Data/Nat/ModEq.lean b/Mathlib/Data/Nat/ModEq.lean index 29ff2fed4da9d..29f36306ea23a 100644 --- a/Mathlib/Data/Nat/ModEq.lean +++ b/Mathlib/Data/Nat/ModEq.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Mathlib.Algebra.Order.Group.Unbundled.Int +import Mathlib.Algebra.Ring.Nat import Mathlib.Data.Int.GCD /-! diff --git a/Mathlib/Data/Nat/Prime/Defs.lean b/Mathlib/Data/Nat/Prime/Defs.lean index d062858f5ac79..30bc5d25b64e8 100644 --- a/Mathlib/Data/Nat/Prime/Defs.lean +++ b/Mathlib/Data/Nat/Prime/Defs.lean @@ -5,8 +5,8 @@ Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Batteries.Data.Nat.Gcd import Mathlib.Algebra.Group.Nat.Units +import Mathlib.Algebra.GroupWithZero.Nat import Mathlib.Algebra.Prime.Defs -import Mathlib.Algebra.Ring.Nat import Mathlib.Data.Nat.Sqrt import Mathlib.Order.Basic @@ -26,9 +26,9 @@ This file deals with prime numbers: natural numbers `p ≥ 2` whose only divisor -/ -open Bool Subtype +assert_not_exists Ring -open Nat +open Bool Subtype Nat namespace Nat variable {n : ℕ} @@ -234,7 +234,7 @@ theorem minFacAux_has_prop {n : ℕ} (n2 : 2 ≤ n) : · exact ⟨k2, dk, a⟩ · refine have := minFac_lemma n k h - minFacAux_has_prop n2 (k + 2) (i + 1) (by simp [k, e, left_distrib, add_right_comm]) + minFacAux_has_prop n2 (k + 2) (i + 1) (by simp [k, e, Nat.left_distrib, add_right_comm]) fun m m2 d => ?_ rcases Nat.eq_or_lt_of_le (a m m2 d) with me | ml · subst me diff --git a/MathlibTest/linear_combination.lean b/MathlibTest/linear_combination.lean index b947d8190146e..96eecab1651c1 100644 --- a/MathlibTest/linear_combination.lean +++ b/MathlibTest/linear_combination.lean @@ -521,7 +521,7 @@ typeclass inference is demanded by the lemmas it orchestrates. This example too (and 73 ms on a good laptop) on an implementation with "minimal" typeclasses everywhere, e.g. lots of `CovariantClass`/`ContravariantClass`, and takes 206 heartbeats (10 ms on a good laptop) on the implementation at the time of joining Mathlib (November 2024). -/ -set_option maxHeartbeats 1100 in +set_option maxHeartbeats 1200 in example {a b : ℝ} (h : a < b) : 0 < b - a := by linear_combination (norm := skip) h exact test_sorry diff --git a/MathlibTest/rewrites.lean b/MathlibTest/rewrites.lean index 2b4fe6758c081..199d5cae1439c 100644 --- a/MathlibTest/rewrites.lean +++ b/MathlibTest/rewrites.lean @@ -1,3 +1,4 @@ +import Mathlib.Algebra.Ring.Nat import Mathlib.Data.Nat.Prime.Defs import Mathlib.CategoryTheory.Category.Basic import Mathlib.Data.List.InsertIdx From 6ab96d745b531402f28209e09cfd6a3a174869f5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 28 Nov 2024 13:58:28 +0000 Subject: [PATCH 373/829] feat: more lemmas about monovarying functions (#19470) From LeanCamCombi --- Mathlib/Order/Monotone/Monovary.lean | 30 ++++++++++++++++++++++++++++ 1 file changed, 30 insertions(+) diff --git a/Mathlib/Order/Monotone/Monovary.lean b/Mathlib/Order/Monotone/Monovary.lean index 9dcfecb90a91f..26d89b2eb6802 100644 --- a/Mathlib/Order/Monotone/Monovary.lean +++ b/Mathlib/Order/Monotone/Monovary.lean @@ -70,6 +70,12 @@ theorem monovaryOn_univ : MonovaryOn f g univ ↔ Monovary f g := theorem antivaryOn_univ : AntivaryOn f g univ ↔ Antivary f g := ⟨fun h _ _ => h trivial trivial, fun h _ _ _ _ hij => h hij⟩ +lemma monovaryOn_iff_monovary : MonovaryOn f g s ↔ Monovary (fun i : s ↦ f i) fun i ↦ g i := by + simp [Monovary, MonovaryOn] + +lemma antivaryOn_iff_antivary : AntivaryOn f g s ↔ Antivary (fun i : s ↦ f i) fun i ↦ g i := by + simp [Antivary, AntivaryOn] + protected theorem MonovaryOn.subset (hst : s ⊆ t) (h : MonovaryOn f g t) : MonovaryOn f g s := fun _ hi _ hj => h (hst hi) (hst hj) @@ -244,6 +250,30 @@ theorem monovaryOn_id_iff : MonovaryOn f id s ↔ MonotoneOn f s := theorem antivaryOn_id_iff : AntivaryOn f id s ↔ AntitoneOn f s := antitoneOn_iff_forall_lt.symm +lemma StrictMono.trans_monovary (hf : StrictMono f) (h : Monovary g f) : Monotone g := + monotone_iff_forall_lt.2 fun _a _b hab ↦ h <| hf hab + +lemma StrictMono.trans_antivary (hf : StrictMono f) (h : Antivary g f) : Antitone g := + antitone_iff_forall_lt.2 fun _a _b hab ↦ h <| hf hab + +lemma StrictAnti.trans_monovary (hf : StrictAnti f) (h : Monovary g f) : Antitone g := + antitone_iff_forall_lt.2 fun _a _b hab ↦ h <| hf hab + +lemma StrictAnti.trans_antivary (hf : StrictAnti f) (h : Antivary g f) : Monotone g := + monotone_iff_forall_lt.2 fun _a _b hab ↦ h <| hf hab + +lemma StrictMonoOn.trans_monovaryOn (hf : StrictMonoOn f s) (h : MonovaryOn g f s) : + MonotoneOn g s := monotoneOn_iff_forall_lt.2 fun _a ha _b hb hab ↦ h ha hb <| hf ha hb hab + +lemma StrictMonoOn.trans_antivaryOn (hf : StrictMonoOn f s) (h : AntivaryOn g f s) : + AntitoneOn g s := antitoneOn_iff_forall_lt.2 fun _a ha _b hb hab ↦ h ha hb <| hf ha hb hab + +lemma StrictAntiOn.trans_monovaryOn (hf : StrictAntiOn f s) (h : MonovaryOn g f s) : + AntitoneOn g s := antitoneOn_iff_forall_lt.2 fun _a ha _b hb hab ↦ h hb ha <| hf ha hb hab + +lemma StrictAntiOn.trans_antivaryOn (hf : StrictAntiOn f s) (h : AntivaryOn g f s) : + MonotoneOn g s := monotoneOn_iff_forall_lt.2 fun _a ha _b hb hab ↦ h hb ha <| hf ha hb hab + end PartialOrder variable [LinearOrder ι] From 5caa22aa9f15c36e2712bf36e5ae0fcf01b2591f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 28 Nov 2024 13:58:30 +0000 Subject: [PATCH 374/829] chore(MetricSpace): don't use `{}` in binders (#19478) The automatic introduction of implicit is very annoying in proofs after having used one of those lemmas. From LeanCamCombi --- Mathlib/Topology/Instances/Rat.lean | 4 ++-- Mathlib/Topology/Instances/Real.lean | 6 ++--- Mathlib/Topology/MetricSpace/Gluing.lean | 4 ++-- Mathlib/Topology/MetricSpace/Pseudo/Defs.lean | 22 +++++++++---------- 4 files changed, 18 insertions(+), 18 deletions(-) diff --git a/Mathlib/Topology/Instances/Rat.lean b/Mathlib/Topology/Instances/Rat.lean index eeb5322f5fccd..a8796eaca78bc 100644 --- a/Mathlib/Topology/Instances/Rat.lean +++ b/Mathlib/Topology/Instances/Rat.lean @@ -96,7 +96,7 @@ theorem uniformContinuous_add : UniformContinuous fun p : ℚ × ℚ => p.1 + p. theorem uniformContinuous_neg : UniformContinuous (@Neg.neg ℚ _) := Metric.uniformContinuous_iff.2 fun ε ε0 => - ⟨_, ε0, fun h => by rw [dist_comm] at h; simpa only [dist_eq, cast_neg, neg_sub_neg] using h⟩ + ⟨_, ε0, fun _ _ h => by simpa only [abs_sub_comm, dist_eq, cast_neg, neg_sub_neg] using h⟩ instance : UniformAddGroup ℚ := UniformAddGroup.mk' Rat.uniformContinuous_add Rat.uniformContinuous_neg @@ -107,7 +107,7 @@ instance : OrderTopology ℚ := induced_orderTopology _ Rat.cast_lt exists_rat_b theorem uniformContinuous_abs : UniformContinuous (abs : ℚ → ℚ) := Metric.uniformContinuous_iff.2 fun ε ε0 => - ⟨ε, ε0, fun h => + ⟨ε, ε0, fun _ _ h => lt_of_le_of_lt (by simpa [Rat.dist_eq] using abs_abs_sub_abs_le_abs_sub _ _) h⟩ instance : TopologicalRing ℚ := inferInstance diff --git a/Mathlib/Topology/Instances/Real.lean b/Mathlib/Topology/Instances/Real.lean index 0cdce97641e0f..522b40e629e4b 100644 --- a/Mathlib/Topology/Instances/Real.lean +++ b/Mathlib/Topology/Instances/Real.lean @@ -33,13 +33,13 @@ instance : NoncompactSpace ℝ := Int.isClosedEmbedding_coe_real.noncompactSpace theorem Real.uniformContinuous_add : UniformContinuous fun p : ℝ × ℝ => p.1 + p.2 := Metric.uniformContinuous_iff.2 fun _ε ε0 => let ⟨δ, δ0, Hδ⟩ := rat_add_continuous_lemma abs ε0 - ⟨δ, δ0, fun h => + ⟨δ, δ0, fun _ _ h => let ⟨h₁, h₂⟩ := max_lt_iff.1 h Hδ h₁ h₂⟩ theorem Real.uniformContinuous_neg : UniformContinuous (@Neg.neg ℝ _) := Metric.uniformContinuous_iff.2 fun ε ε0 => - ⟨_, ε0, fun h => by rw [dist_comm] at h; simpa only [Real.dist_eq, neg_sub_neg] using h⟩ + ⟨_, ε0, fun _ _ h => by simpa only [abs_sub_comm, Real.dist_eq, neg_sub_neg] using h⟩ instance : ContinuousStar ℝ := ⟨continuous_id⟩ @@ -97,7 +97,7 @@ theorem Real.uniformContinuous_inv (s : Set ℝ) {r : ℝ} (r0 : 0 < r) (H : ∀ theorem Real.uniformContinuous_abs : UniformContinuous (abs : ℝ → ℝ) := Metric.uniformContinuous_iff.2 fun ε ε0 => - ⟨ε, ε0, lt_of_le_of_lt (abs_abs_sub_abs_le_abs_sub _ _)⟩ + ⟨ε, ε0, fun _ _ ↦ lt_of_le_of_lt (abs_abs_sub_abs_le_abs_sub _ _)⟩ theorem Real.continuous_inv : Continuous fun a : { r : ℝ // r ≠ 0 } => a.val⁻¹ := continuousOn_inv₀.restrict diff --git a/Mathlib/Topology/MetricSpace/Gluing.lean b/Mathlib/Topology/MetricSpace/Gluing.lean index f11b797b3559e..e33a9b03840db 100644 --- a/Mathlib/Topology/MetricSpace/Gluing.lean +++ b/Mathlib/Topology/MetricSpace/Gluing.lean @@ -167,7 +167,7 @@ theorem Sum.mem_uniformity_iff_glueDist (hε : 0 < ε) (s : Set ((X ⊕ Y) × (X · exact absurd h.2 (le_glueDist_inr_inl _ _ _ _ _).not_lt · exact hY h.1.2 · rintro ⟨ε, ε0, H⟩ - constructor <;> exact ⟨ε, ε0, fun h => H _ _ h⟩ + constructor <;> exact ⟨ε, ε0, fun _ _ h => H _ _ h⟩ /-- Given two maps `Φ` and `Ψ` intro metric spaces `X` and `Y` such that the distances between `Φ p` and `Φ q`, and between `Ψ p` and `Ψ q`, coincide up to `2 ε` where `ε > 0`, one can almost @@ -243,7 +243,7 @@ private theorem Sum.mem_uniformity (s : Set ((X ⊕ Y) × (X ⊕ Y))) : · cases not_le_of_lt (lt_of_lt_of_le h (min_le_right _ _)) Sum.one_le_dist_inr_inl · exact hY (lt_of_lt_of_le h (le_trans (min_le_left _ _) (min_le_right _ _))) · rintro ⟨ε, ε0, H⟩ - constructor <;> rw [Filter.mem_map, mem_uniformity_dist] <;> exact ⟨ε, ε0, fun h => H _ _ h⟩ + constructor <;> rw [Filter.mem_map, mem_uniformity_dist] <;> exact ⟨ε, ε0, fun _ _ h => H _ _ h⟩ /-- The distance on the disjoint union indeed defines a metric space. All the distance properties follow from our choice of the distance. The harder work is to show that the uniform structure diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean b/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean index 6f5392e44495c..13165e067a52c 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean @@ -671,15 +671,15 @@ theorem uniformity_basis_dist_le_pow {r : ℝ} (h0 : 0 < r) (h1 : r < 1) : ⟨n, trivial, hn.le⟩ theorem mem_uniformity_dist {s : Set (α × α)} : - s ∈ 𝓤 α ↔ ∃ ε > 0, ∀ {a b : α}, dist a b < ε → (a, b) ∈ s := + s ∈ 𝓤 α ↔ ∃ ε > 0, ∀ ⦃a b : α⦄, dist a b < ε → (a, b) ∈ s := uniformity_basis_dist.mem_uniformity_iff /-- A constant size neighborhood of the diagonal is an entourage. -/ theorem dist_mem_uniformity {ε : ℝ} (ε0 : 0 < ε) : { p : α × α | dist p.1 p.2 < ε } ∈ 𝓤 α := - mem_uniformity_dist.2 ⟨ε, ε0, id⟩ + mem_uniformity_dist.2 ⟨ε, ε0, fun _ _ ↦ id⟩ theorem uniformContinuous_iff [PseudoMetricSpace β] {f : α → β} : - UniformContinuous f ↔ ∀ ε > 0, ∃ δ > 0, ∀ {a b : α}, dist a b < δ → dist (f a) (f b) < ε := + UniformContinuous f ↔ ∀ ε > 0, ∃ δ > 0, ∀ ⦃a b : α⦄, dist a b < δ → dist (f a) (f b) < ε := uniformity_basis_dist.uniformContinuous_iff uniformity_basis_dist theorem uniformContinuousOn_iff [PseudoMetricSpace β] {f : α → β} {s : Set α} : @@ -710,7 +710,7 @@ theorem eventually_nhds_iff_ball {p : α → Prop} : in a pseudo-metric space. -/ theorem eventually_nhds_prod_iff {f : Filter ι} {x₀ : α} {p : α × ι → Prop} : (∀ᶠ x in 𝓝 x₀ ×ˢ f, p x) ↔ ∃ ε > (0 : ℝ), ∃ pa : ι → Prop, (∀ᶠ i in f, pa i) ∧ - ∀ {x}, dist x x₀ < ε → ∀ {i}, pa i → p (x, i) := by + ∀ ⦃x⦄, dist x x₀ < ε → ∀ ⦃i⦄, pa i → p (x, i) := by refine (nhds_basis_ball.prod f.basis_sets).eventually_iff.trans ?_ simp only [Prod.exists, forall_prod_set, id, mem_ball, and_assoc, exists_and_left, and_imp] rfl @@ -719,11 +719,11 @@ theorem eventually_nhds_prod_iff {f : Filter ι} {x₀ : α} {p : α × ι → P in a pseudo-metric space. -/ theorem eventually_prod_nhds_iff {f : Filter ι} {x₀ : α} {p : ι × α → Prop} : (∀ᶠ x in f ×ˢ 𝓝 x₀, p x) ↔ ∃ pa : ι → Prop, (∀ᶠ i in f, pa i) ∧ - ∃ ε > 0, ∀ {i}, pa i → ∀ {x}, dist x x₀ < ε → p (i, x) := by + ∃ ε > 0, ∀ ⦃i⦄, pa i → ∀ ⦃x⦄, dist x x₀ < ε → p (i, x) := by rw [eventually_swap_iff, Metric.eventually_nhds_prod_iff] constructor <;> · rintro ⟨a1, a2, a3, a4, a5⟩ - exact ⟨a3, a4, a1, a2, fun b1 b2 b3 => a5 b3 b1⟩ + exact ⟨a3, a4, a1, a2, fun _ b1 b2 b3 => a5 b3 b1⟩ theorem nhds_basis_closedBall : (𝓝 x).HasBasis (fun ε : ℝ => 0 < ε) (closedBall x) := nhds_basis_uniformity uniformity_basis_dist_le @@ -768,27 +768,27 @@ theorem mem_nhdsWithin_iff {t : Set α} : s ∈ 𝓝[t] x ↔ ∃ ε > 0, ball x theorem tendsto_nhdsWithin_nhdsWithin [PseudoMetricSpace β] {t : Set β} {f : α → β} {a b} : Tendsto f (𝓝[s] a) (𝓝[t] b) ↔ - ∀ ε > 0, ∃ δ > 0, ∀ {x : α}, x ∈ s → dist x a < δ → f x ∈ t ∧ dist (f x) b < ε := + ∀ ε > 0, ∃ δ > 0, ∀ ⦃x : α⦄, x ∈ s → dist x a < δ → f x ∈ t ∧ dist (f x) b < ε := (nhdsWithin_basis_ball.tendsto_iff nhdsWithin_basis_ball).trans <| by simp only [inter_comm _ s, inter_comm _ t, mem_inter_iff, and_imp, gt_iff_lt, mem_ball] theorem tendsto_nhdsWithin_nhds [PseudoMetricSpace β] {f : α → β} {a b} : Tendsto f (𝓝[s] a) (𝓝 b) ↔ - ∀ ε > 0, ∃ δ > 0, ∀ {x : α}, x ∈ s → dist x a < δ → dist (f x) b < ε := by + ∀ ε > 0, ∃ δ > 0, ∀ ⦃x : α⦄, x ∈ s → dist x a < δ → dist (f x) b < ε := by rw [← nhdsWithin_univ b, tendsto_nhdsWithin_nhdsWithin] simp only [mem_univ, true_and] theorem tendsto_nhds_nhds [PseudoMetricSpace β] {f : α → β} {a b} : - Tendsto f (𝓝 a) (𝓝 b) ↔ ∀ ε > 0, ∃ δ > 0, ∀ {x : α}, dist x a < δ → dist (f x) b < ε := + Tendsto f (𝓝 a) (𝓝 b) ↔ ∀ ε > 0, ∃ δ > 0, ∀ ⦃x : α⦄, dist x a < δ → dist (f x) b < ε := nhds_basis_ball.tendsto_iff nhds_basis_ball theorem continuousAt_iff [PseudoMetricSpace β] {f : α → β} {a : α} : - ContinuousAt f a ↔ ∀ ε > 0, ∃ δ > 0, ∀ {x : α}, dist x a < δ → dist (f x) (f a) < ε := by + ContinuousAt f a ↔ ∀ ε > 0, ∃ δ > 0, ∀ ⦃x : α⦄, dist x a < δ → dist (f x) (f a) < ε := by rw [ContinuousAt, tendsto_nhds_nhds] theorem continuousWithinAt_iff [PseudoMetricSpace β] {f : α → β} {a : α} {s : Set α} : ContinuousWithinAt f s a ↔ - ∀ ε > 0, ∃ δ > 0, ∀ {x : α}, x ∈ s → dist x a < δ → dist (f x) (f a) < ε := by + ∀ ε > 0, ∃ δ > 0, ∀ ⦃x : α⦄, x ∈ s → dist x a < δ → dist (f x) (f a) < ε := by rw [ContinuousWithinAt, tendsto_nhdsWithin_nhds] theorem continuousOn_iff [PseudoMetricSpace β] {f : α → β} {s : Set α} : From bbf85a5ee0f4086b3da87736c052cd02fc0c9319 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Thu, 28 Nov 2024 13:58:31 +0000 Subject: [PATCH 375/829] feat(MeasureTheory): add Measure.fst_add, Measure.fst_sum (#19564) and similar lemmas for `snd`. --- Mathlib/MeasureTheory/Measure/Prod.lean | 22 ++++++++++++++++++++++ 1 file changed, 22 insertions(+) diff --git a/Mathlib/MeasureTheory/Measure/Prod.lean b/Mathlib/MeasureTheory/Measure/Prod.lean index e82f876ba5961..4a67f092f79cf 100644 --- a/Mathlib/MeasureTheory/Measure/Prod.lean +++ b/Mathlib/MeasureTheory/Measure/Prod.lean @@ -874,6 +874,17 @@ theorem fst_map_prod_mk {X : α → β} {Y : α → γ} {μ : Measure α} (hY : Measurable Y) : (μ.map fun a => (X a, Y a)).fst = μ.map X := fst_map_prod_mk₀ hY.aemeasurable +@[simp] +lemma fst_add {μ ν : Measure (α × β)} : (μ + ν).fst = μ.fst + ν.fst := by + ext s hs + simp_rw [coe_add, Pi.add_apply, fst_apply hs, coe_add, Pi.add_apply] + +lemma fst_sum {ι : Type*} (μ : ι → Measure (α × β)) : (sum μ).fst = sum (fun n ↦ (μ n).fst) := by + ext s hs + rw [fst_apply hs, sum_apply, sum_apply _ hs] + · simp_rw [fst_apply hs] + · exact measurable_fst hs + @[gcongr] theorem fst_mono {μ : Measure (α × β)} (h : ρ ≤ μ) : ρ.fst ≤ μ.fst := map_mono h measurable_fst @@ -921,6 +932,17 @@ theorem snd_map_prod_mk {X : α → β} {Y : α → γ} {μ : Measure α} (hX : (μ.map fun a => (X a, Y a)).snd = μ.map Y := snd_map_prod_mk₀ hX.aemeasurable +@[simp] +lemma snd_add {μ ν : Measure (α × β)} : (μ + ν).snd = μ.snd + ν.snd := by + ext s hs + simp_rw [coe_add, Pi.add_apply, snd_apply hs, coe_add, Pi.add_apply] + +lemma snd_sum {ι : Type*} (μ : ι → Measure (α × β)) : (sum μ).snd = sum (fun n ↦ (μ n).snd) := by + ext s hs + rw [snd_apply hs, sum_apply, sum_apply _ hs] + · simp_rw [snd_apply hs] + · exact measurable_snd hs + @[gcongr] theorem snd_mono {μ : Measure (α × β)} (h : ρ ≤ μ) : ρ.snd ≤ μ.snd := map_mono h measurable_snd From 2b8bedc5045b52d40504db23dd34036eb6ae4e46 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Thu, 28 Nov 2024 14:24:30 +0000 Subject: [PATCH 376/829] feat(Analysis/NormedSpace/FunctionSeries): add eventually versions (#16543) Co-authored-by: Chris Birkbeck --- .../Analysis/NormedSpace/FunctionSeries.lean | 37 +++++++++++++++++++ 1 file changed, 37 insertions(+) diff --git a/Mathlib/Analysis/NormedSpace/FunctionSeries.lean b/Mathlib/Analysis/NormedSpace/FunctionSeries.lean index d4783a6874149..3f07c20513c97 100644 --- a/Mathlib/Analysis/NormedSpace/FunctionSeries.lean +++ b/Mathlib/Analysis/NormedSpace/FunctionSeries.lean @@ -44,6 +44,34 @@ theorem tendstoUniformlyOn_tsum_nat {f : ℕ → β → F} {u : ℕ → ℝ} (hu s := fun v hv => tendsto_finset_range.eventually (tendstoUniformlyOn_tsum hu hfu v hv) +/-- An infinite sum of functions with eventually summable sup norm is the uniform limit of its +partial sums. Version relative to a set, with general index set. -/ +theorem tendstoUniformlyOn_tsum_of_cofinite_eventually {ι : Type*} {f : ι → β → F} {u : ι → ℝ} + (hu : Summable u) {s : Set β} (hfu : ∀ᶠ n in cofinite, ∀ x ∈ s, ‖f n x‖ ≤ u n) : + TendstoUniformlyOn (fun t x => ∑ n ∈ t, f n x) (fun x => ∑' n, f n x) atTop s := by + classical + refine tendstoUniformlyOn_iff.2 fun ε εpos => ?_ + have := (tendsto_order.1 (tendsto_tsum_compl_atTop_zero u)).2 _ εpos + simp only [not_forall, Classical.not_imp, not_le, gt_iff_lt, + eventually_atTop, ge_iff_le, Finset.le_eq_subset] at * + obtain ⟨t, ht⟩ := this + rw [eventually_iff_exists_mem] at hfu + obtain ⟨N, hN, HN⟩ := hfu + refine ⟨hN.toFinset ∪ t, fun n hn x hx => ?_⟩ + have A : Summable fun n => ‖f n x‖ := by + apply Summable.add_compl (s := hN.toFinset) Summable.of_finite + apply Summable.of_nonneg_of_le (fun _ ↦ norm_nonneg _) _ (hu.subtype _) + simp only [comp_apply, Subtype.forall, Set.mem_compl_iff, Finset.mem_coe] + aesop + rw [dist_eq_norm, ← sum_add_tsum_subtype_compl A.of_norm n, add_sub_cancel_left] + apply lt_of_le_of_lt _ (ht n (Finset.union_subset_right hn)) + apply (norm_tsum_le_tsum_norm (A.subtype _)).trans + apply tsum_le_tsum _ (A.subtype _) (hu.subtype _) + simp only [comp_apply, Subtype.forall, imp_false] + apply fun i hi => HN i ?_ x hx + have : ¬ i ∈ hN.toFinset := fun hg ↦ hi (Finset.union_subset_left hn hg) + aesop + /-- An infinite sum of functions with summable sup norm is the uniform limit of its partial sums. Version with general index set. -/ theorem tendstoUniformly_tsum {f : α → β → F} (hu : Summable u) (hfu : ∀ n x, ‖f n x‖ ≤ u n) : @@ -59,6 +87,15 @@ theorem tendstoUniformly_tsum_nat {f : ℕ → β → F} {u : ℕ → ℝ} (hu : atTop := fun v hv => tendsto_finset_range.eventually (tendstoUniformly_tsum hu hfu v hv) +/-- An infinite sum of functions with eventually summable sup norm is the uniform limit of its +partial sums. Version with general index set. -/ +theorem tendstoUniformly_tsum_of_cofinite_eventually {ι : Type*} {f : ι → β → F} {u : ι → ℝ} + (hu : Summable u) (hfu : ∀ᶠ (n : ι) in cofinite, ∀ x : β, ‖f n x‖ ≤ u n) : + TendstoUniformly (fun t x => ∑ n ∈ t, f n x) (fun x => ∑' n, f n x) atTop := by + rw [← tendstoUniformlyOn_univ] + apply tendstoUniformlyOn_tsum_of_cofinite_eventually hu + simpa using hfu + /-- An infinite sum of functions with summable sup norm is continuous on a set if each individual function is. -/ theorem continuousOn_tsum [TopologicalSpace β] {f : α → β → F} {s : Set β} From d8c4e54c741027c4114fcff4c121fe3ce9df5dd2 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Thu, 28 Nov 2024 14:24:31 +0000 Subject: [PATCH 377/829] feat(GroupTheory/Subgroup/Centralizer): The `N/C` theorem (#19006) This PR adds the `N/C` theorem that `Normalizer(H)/Centralizer(H)` embeds as a subgroup of `Aut(H)`. --- Mathlib/GroupTheory/Subgroup/Centralizer.lean | 20 +++++++++++++++++++ 1 file changed, 20 insertions(+) diff --git a/Mathlib/GroupTheory/Subgroup/Centralizer.lean b/Mathlib/GroupTheory/Subgroup/Centralizer.lean index 664c66c389c5c..9a0fb7f77947e 100644 --- a/Mathlib/GroupTheory/Subgroup/Centralizer.lean +++ b/Mathlib/GroupTheory/Subgroup/Centralizer.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Kexing Ying. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kexing Ying -/ +import Mathlib.Algebra.GroupWithZero.Action.Basic import Mathlib.GroupTheory.Subgroup.Center import Mathlib.GroupTheory.Submonoid.Centralizer @@ -94,4 +95,23 @@ abbrev closureCommGroupOfComm {k : Set G} (hcomm : ∀ x ∈ k, ∀ y ∈ k, x * have := closure_le_centralizer_centralizer k Subtype.ext <| Set.centralizer_centralizer_comm_of_comm hcomm _ (this h₁) _ (this h₂) } +/-- The conjugation action of N(H) on H. -/ +@[simps] +instance : MulDistribMulAction H.normalizer H where + smul g h := ⟨g * h * g⁻¹, (g.2 h).mp h.2⟩ + one_smul g := by simp [HSMul.hSMul] + mul_smul := by simp [HSMul.hSMul, mul_assoc] + smul_one := by simp [HSMul.hSMul] + smul_mul := by simp [HSMul.hSMul] + +/-- The homomorphism N(H) → Aut(H) with kernel C(H). -/ +@[simps!] +def normalizerMonoidHom : H.normalizer →* MulAut H := + MulDistribMulAction.toMulAut H.normalizer H + +theorem normalizerMonoidHom_ker : + H.normalizerMonoidHom.ker = (Subgroup.centralizer H).subgroupOf H.normalizer := by + simp [Subgroup.ext_iff, DFunLike.ext_iff, Subtype.ext_iff, + mem_subgroupOf, mem_centralizer_iff, eq_mul_inv_iff_mul_eq, eq_comm] + end Subgroup From 4b375340f6cdf5c73367b401c83b834075a93e47 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Thu, 28 Nov 2024 14:24:33 +0000 Subject: [PATCH 378/829] feat(Algebra/Category/ModuleCat/Products): show that arbitrary coproducts agree with direct sum (#19024) --- .../Algebra/Category/ModuleCat/Products.lean | 50 +++++++++++++++++++ 1 file changed, 50 insertions(+) diff --git a/Mathlib/Algebra/Category/ModuleCat/Products.lean b/Mathlib/Algebra/Category/ModuleCat/Products.lean index b986cea15a7f5..9b504d0c1dbe0 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Products.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Products.lean @@ -5,6 +5,7 @@ Authors: Kim Morrison -/ import Mathlib.Algebra.Category.ModuleCat.Basic import Mathlib.LinearAlgebra.Pi +import Mathlib.Algebra.DirectSum.Module import Mathlib.Tactic.CategoryTheory.Elementwise /-! @@ -23,6 +24,7 @@ namespace ModuleCat variable {R : Type u} [Ring R] variable {ι : Type v} (Z : ι → ModuleCatMax.{v, w} R) +section product /-- The product cone induced by the concrete product. -/ def productCone : Fan Z := @@ -60,4 +62,52 @@ theorem piIsoPi_hom_ker_subtype (i : ι) : (piIsoPi Z).hom ≫ (LinearMap.proj i : (∀ i : ι, Z i) →ₗ[R] Z i) = Pi.π Z i := IsLimit.conePointUniqueUpToIso_inv_comp _ (limit.isLimit _) (Discrete.mk i) +end product + +section coproduct + +open DirectSum + +variable [DecidableEq ι] + +/-- The coproduct cone induced by the concrete product. -/ +def coproductCocone : Cofan Z := + Cofan.mk (ModuleCat.of R (⨁ i : ι, Z i)) fun i => (DirectSum.lof R ι (fun i ↦ Z i) i) + +/-- The concrete coproduct cone is limiting. -/ +def coproductCoconeIsColimit : IsColimit (coproductCocone Z) where + desc s := DirectSum.toModule R ι _ fun i ↦ s.ι.app ⟨i⟩ + fac := by + rintro s ⟨i⟩ + ext (x : Z i) + simpa only [Discrete.functor_obj_eq_as, coproductCocone, Cofan.mk_pt, Functor.const_obj_obj, + Cofan.mk_ι_app, coe_comp, Function.comp_apply] using + DirectSum.toModule_lof (ι := ι) R (M := fun i ↦ Z i) i x + uniq := by + rintro s f h + refine DirectSum.linearMap_ext _ fun i ↦ ?_ + ext x + simpa only [LinearMap.coe_comp, Function.comp_apply, toModule_lof] using + congr($(h ⟨i⟩) x) + +variable [HasCoproduct Z] + +/-- The categorical coproduct of a family of objects in `ModuleCat` +agrees with direct sum. +-/ +noncomputable def coprodIsoDirectSum : ∐ Z ≅ ModuleCat.of R (⨁ i, Z i) := + colimit.isoColimitCocone ⟨_, coproductCoconeIsColimit Z⟩ + +@[simp, elementwise] +theorem ι_coprodIsoDirectSum_hom (i : ι) : + Sigma.ι Z i ≫ (coprodIsoDirectSum Z).hom = DirectSum.lof R ι (fun i ↦ Z i) i := + colimit.isoColimitCocone_ι_hom _ _ + +@[simp, elementwise] +theorem lof_coprodIsoDirectSum_inv (i : ι) : + DirectSum.lof R ι (fun i ↦ Z i) i ≫ (coprodIsoDirectSum Z).inv = Sigma.ι Z i := + (coproductCoconeIsColimit Z).comp_coconePointUniqueUpToIso_hom (colimit.isColimit _) _ + +end coproduct + end ModuleCat From 061f06f308d94a2142f19b50ed742690ae175afa Mon Sep 17 00:00:00 2001 From: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com> Date: Thu, 28 Nov 2024 14:24:34 +0000 Subject: [PATCH 379/829] feat: Polynomial.isRoot_of_isRoot_iff_dvd_derivative_mul (#19146) Co-authored-by: ChrisHughes24 --- Mathlib/Algebra/Polynomial/FieldDivision.lean | 24 +++++++++++++++++ .../IsAlgClosed/AlgebraicClosure.lean | 27 +++++++++++++++++++ 2 files changed, 51 insertions(+) diff --git a/Mathlib/Algebra/Polynomial/FieldDivision.lean b/Mathlib/Algebra/Polynomial/FieldDivision.lean index 0223ce382a22e..bc0f49ad785ba 100644 --- a/Mathlib/Algebra/Polynomial/FieldDivision.lean +++ b/Mathlib/Algebra/Polynomial/FieldDivision.lean @@ -168,6 +168,30 @@ theorem lt_rootMultiplicity_iff_isRoot_iterate_derivative ⟨fun hn _ hm ↦ isRoot_iterate_derivative_of_lt_rootMultiplicity <| Nat.lt_of_le_of_lt hm hn, fun hr ↦ lt_rootMultiplicity_of_isRoot_iterate_derivative h hr⟩ +/-- A sufficient condition for the set of roots of a nonzero polynomial `f` to be a subset of the +set of roots of `g` is that `f` divides `f.derivative * g`. Over an algebraically closed field of +characteristic zero, this is also a necessary condition. +See `isRoot_of_isRoot_iff_dvd_derivative_mul` -/ +theorem isRoot_of_isRoot_of_dvd_derivative_mul [CharZero R] {f g : R[X]} (hf0 : f ≠ 0) + (hfd : f ∣ f.derivative * g) {a : R} (haf : f.IsRoot a) : g.IsRoot a := by + rcases hfd with ⟨r, hr⟩ + by_cases hdf0 : derivative f = 0 + · rw [eq_C_of_derivative_eq_zero hdf0] at haf hf0 + simp only [eval_C, derivative_C, zero_mul, dvd_zero, iff_true, IsRoot.def] at haf + rw [haf, map_zero] at hf0 + exact (hf0 rfl).elim + by_contra hg + have hdfg0 : f.derivative * g ≠ 0 := mul_ne_zero hdf0 (by rintro rfl; simp at hg) + have hr' := congr_arg (rootMultiplicity a) hr + rw [rootMultiplicity_mul hdfg0, derivative_rootMultiplicity_of_root haf, + rootMultiplicity_eq_zero hg, add_zero, rootMultiplicity_mul (hr ▸ hdfg0), add_comm, + Nat.sub_eq_iff_eq_add (Nat.succ_le_iff.2 ((rootMultiplicity_pos hf0).2 haf))] at hr' + refine lt_irrefl (rootMultiplicity a f) ?_ + refine lt_of_lt_of_le (Nat.lt_succ_self _) + (le_trans (le_add_of_nonneg_left (Nat.zero_le (rootMultiplicity a r))) ?_) + conv_rhs => rw [hr'] + simp [add_assoc] + section NormalizationMonoid variable [NormalizationMonoid R] diff --git a/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean b/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean index eac4c949078d3..d91f6a62d0512 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean @@ -411,3 +411,30 @@ instance {L : Type*} [Field k] [Field L] [Algebra k L] [Algebra.IsAlgebraic k L] isAlgClosed := inferInstance end AlgebraicClosure + +/-- Over an algebraically closed field of characteristic zero a necessary and sufficient condition +for the set of roots of a nonzero polynomial `f` to be a subset of the set of roots of `g` is that +`f` divides `f.derivative * g`. Over an integral domain, this is a sufficient but not necessary +condition. See `isRoot_of_isRoot_of_dvd_derivative_mul` -/ +theorem Polynomial.isRoot_of_isRoot_iff_dvd_derivative_mul {K : Type*} [Field K] + [IsAlgClosed K] [CharZero K] {f g : K[X]} (hf0 : f ≠ 0) : + (∀ x, IsRoot f x → IsRoot g x) ↔ f ∣ f.derivative * g := by + refine ⟨?_, isRoot_of_isRoot_of_dvd_derivative_mul hf0⟩ + by_cases hg0 : g = 0 + · simp [hg0] + by_cases hdf0 : derivative f = 0 + · rw [eq_C_of_derivative_eq_zero hdf0] + simp only [eval_C, derivative_C, zero_mul, dvd_zero, implies_true] + have hdg : f.derivative * g ≠ 0 := mul_ne_zero hdf0 hg0 + classical rw [Splits.dvd_iff_roots_le_roots (IsAlgClosed.splits f) hf0 hdg, Multiset.le_iff_count] + simp only [count_roots, rootMultiplicity_mul hdg] + refine forall_imp fun a => ?_ + by_cases haf : f.eval a = 0 + · have h0 : 0 < f.rootMultiplicity a := (rootMultiplicity_pos hf0).2 haf + rw [derivative_rootMultiplicity_of_root haf] + intro h + calc rootMultiplicity a f + = rootMultiplicity a f - 1 + 1 := (Nat.sub_add_cancel (Nat.succ_le_iff.1 h0)).symm + _ ≤ rootMultiplicity a f - 1 + rootMultiplicity a g := add_le_add le_rfl (Nat.succ_le_iff.1 + ((rootMultiplicity_pos hg0).2 (h haf))) + · simp [haf, rootMultiplicity_eq_zero haf] From 28d3cb713ce95bc80e8e0c4c8ca0122dde26294b Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Thu, 28 Nov 2024 14:24:35 +0000 Subject: [PATCH 380/829] feat(GroupTheory/SpecificGroups/Cyclic): Cardinality of automorphism group (#19180) This PR computes the cardinality of the automorphism group of a cyclic group. This is useful when paired with the N/C theorem in PR #19006. I added a new file to avoid adding extra imports to `Data/ZMod/Basic.lean`. --- Mathlib.lean | 1 + Mathlib/Data/ZMod/Aut.lean | 28 +++++++++++++++++++ .../GroupTheory/SpecificGroups/Cyclic.lean | 8 ++++++ 3 files changed, 37 insertions(+) create mode 100644 Mathlib/Data/ZMod/Aut.lean diff --git a/Mathlib.lean b/Mathlib.lean index 742aee5ebc6c8..48ff47a292ad9 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2901,6 +2901,7 @@ import Mathlib.Data.Vector3 import Mathlib.Data.W.Basic import Mathlib.Data.W.Cardinal import Mathlib.Data.W.Constructions +import Mathlib.Data.ZMod.Aut import Mathlib.Data.ZMod.Basic import Mathlib.Data.ZMod.Coprime import Mathlib.Data.ZMod.Defs diff --git a/Mathlib/Data/ZMod/Aut.lean b/Mathlib/Data/ZMod/Aut.lean new file mode 100644 index 0000000000000..88711c2b14b8a --- /dev/null +++ b/Mathlib/Data/ZMod/Aut.lean @@ -0,0 +1,28 @@ +/- +Copyright (c) 2024 Thomas Browning. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Thomas Browning +-/ +import Mathlib.Algebra.Ring.AddAut +import Mathlib.Data.ZMod.Basic + +/-! +# Automorphism Group of `ZMod`. +-/ + +namespace ZMod + +variable (n : ℕ) + +/-- The automorphism group of `ZMod n` is isomorphic to the group of units of `ZMod n`. -/ +@[simps] +def AddAutEquivUnits : AddAut (ZMod n) ≃* (ZMod n)ˣ := + have h (f : AddAut (ZMod n)) (x : ZMod n) : f 1 * x = f x := by + rw [mul_comm, ← x.intCast_zmod_cast, ← zsmul_eq_mul, ← map_zsmul, zsmul_one] + { toFun := fun f ↦ Units.mkOfMulEqOne (f 1) (f⁻¹ 1) ((h f _).trans (f.inv_apply_self _ _)) + invFun := AddAut.mulLeft + left_inv := fun f ↦ by simp [DFunLike.ext_iff, Units.smul_def, h] + right_inv := fun x ↦ by simp [Units.ext_iff, Units.smul_def] + map_mul' := fun f g ↦ by simp [Units.ext_iff, h] } + +end ZMod diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index d13254a943cc7..fd2f7b649dab0 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl -/ import Mathlib.Algebra.Order.BigOperators.Ring.Finset import Mathlib.Data.Nat.Totient +import Mathlib.Data.ZMod.Aut import Mathlib.Data.ZMod.Quotient import Mathlib.GroupTheory.OrderOfElement import Mathlib.GroupTheory.Subgroup.Simple @@ -724,6 +725,13 @@ noncomputable def mulEquivOfPrimeCardEq {p : ℕ} [Group G] [Group G'] apply mulEquivOfCyclicCardEq exact hG.trans hH.symm +theorem IsCyclic.card_mulAut [Group G] [Finite G] [h : IsCyclic G] : + Nat.card (MulAut G) = Nat.totient (Nat.card G) := by + have : NeZero (Nat.card G) := ⟨Nat.card_pos.ne'⟩ + rw [← ZMod.card_units_eq_totient, ← Nat.card_eq_fintype_card] + exact Nat.card_congr ((MulAut.congr (zmodCyclicMulEquiv h)).toEquiv.symm.trans + (MulEquiv.toAdditive.trans (ZMod.AddAutEquivUnits (Nat.card G)).toEquiv)) + end ZMod section generator From 0742f6cb7259675f3110e514d9d186369407ca06 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 28 Nov 2024 14:24:37 +0000 Subject: [PATCH 381/829] feat: interaction of `Finset.sup` and `Nat.cast` (#19245) Co-authored-by: Andrew Yang --- Mathlib.lean | 1 + Mathlib/Algebra/Order/Ring/Finset.lean | 39 ++++++++++++++++++++++++++ Mathlib/Data/Nat/Cast/Order/Ring.lean | 4 +-- 3 files changed, 42 insertions(+), 2 deletions(-) create mode 100644 Mathlib/Algebra/Order/Ring/Finset.lean diff --git a/Mathlib.lean b/Mathlib.lean index 48ff47a292ad9..4ace0a7b6e35c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -743,6 +743,7 @@ import Mathlib.Algebra.Order.Ring.Canonical import Mathlib.Algebra.Order.Ring.Cast import Mathlib.Algebra.Order.Ring.Cone import Mathlib.Algebra.Order.Ring.Defs +import Mathlib.Algebra.Order.Ring.Finset import Mathlib.Algebra.Order.Ring.InjSurj import Mathlib.Algebra.Order.Ring.Int import Mathlib.Algebra.Order.Ring.Nat diff --git a/Mathlib/Algebra/Order/Ring/Finset.lean b/Mathlib/Algebra/Order/Ring/Finset.lean new file mode 100644 index 0000000000000..0b92b63875cd3 --- /dev/null +++ b/Mathlib/Algebra/Order/Ring/Finset.lean @@ -0,0 +1,39 @@ +/- +Copyright (c) 2022 Eric Wieser, Yaël Dillies, Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser, Yaël Dillies, Andrew Yang +-/ +import Mathlib.Algebra.Order.Field.Canonical.Defs +import Mathlib.Data.Finset.Lattice.Fold +import Mathlib.Data.Nat.Cast.Order.Ring + +/-! +# `Finset.sup` of `Nat.cast` +-/ + +open Finset + +namespace Nat +variable {ι R : Type*} + +section LinearOrderedSemiring +variable [LinearOrderedSemiring R] {s : Finset ι} + +set_option linter.docPrime false in +@[simp, norm_cast] +lemma cast_finsetSup' (f : ι → ℕ) (hs) : ((s.sup' hs f : ℕ) : R) = s.sup' hs fun i ↦ (f i : R) := + comp_sup'_eq_sup'_comp _ _ cast_max + +set_option linter.docPrime false in +@[simp, norm_cast] +lemma cast_finsetInf' (f : ι → ℕ) (hs) : (↑(s.inf' hs f) : R) = s.inf' hs fun i ↦ (f i : R) := + comp_inf'_eq_inf'_comp _ _ cast_min + +end LinearOrderedSemiring + +@[simp, norm_cast] +lemma cast_finsetSup [CanonicallyLinearOrderedSemifield R] (s : Finset ι) (f : ι → ℕ) : + (↑(s.sup f) : R) = s.sup fun i ↦ (f i : R) := + comp_sup_eq_sup_comp _ cast_max (by simp) + +end Nat diff --git a/Mathlib/Data/Nat/Cast/Order/Ring.lean b/Mathlib/Data/Nat/Cast/Order/Ring.lean index 97687710d7557..693d3ef4c6797 100644 --- a/Mathlib/Data/Nat/Cast/Order/Ring.lean +++ b/Mathlib/Data/Nat/Cast/Order/Ring.lean @@ -37,11 +37,11 @@ theorem ofNat_nonneg {α} [OrderedSemiring α] (n : ℕ) [n.AtLeastTwo] : ofNat_nonneg' n @[simp, norm_cast] -theorem cast_min {α} [LinearOrderedSemiring α] {a b : ℕ} : ((min a b : ℕ) : α) = min (a : α) b := +theorem cast_min {α} [LinearOrderedSemiring α] (m n : ℕ) : (↑(min m n : ℕ) : α) = min (m : α) n := (@mono_cast α _).map_min @[simp, norm_cast] -theorem cast_max {α} [LinearOrderedSemiring α] {a b : ℕ} : ((max a b : ℕ) : α) = max (a : α) b := +theorem cast_max {α} [LinearOrderedSemiring α] (m n : ℕ) : (↑(max m n : ℕ) : α) = max (m : α) n := (@mono_cast α _).map_max section Nontrivial From cbd4093394f209fd38ef7c0a95acc251f0e26554 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 28 Nov 2024 14:24:38 +0000 Subject: [PATCH 382/829] =?UTF-8?q?feat:=20`{a,=20b}=20*=20s=20=3D=20a=20?= =?UTF-8?q?=E2=80=A2=20s=20=E2=88=AA=20b=20=E2=80=A2=20s`=20(#19406)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit and other simple pointwise lemmas. From GrowthInGroups (LeanCamCombi) --- Mathlib/Data/Set/Pointwise/SMul.lean | 50 ++++++++++++++++++++++++---- 1 file changed, 44 insertions(+), 6 deletions(-) diff --git a/Mathlib/Data/Set/Pointwise/SMul.lean b/Mathlib/Data/Set/Pointwise/SMul.lean index d37ecd85db188..4a60e0e49e98f 100644 --- a/Mathlib/Data/Set/Pointwise/SMul.lean +++ b/Mathlib/Data/Set/Pointwise/SMul.lean @@ -3,13 +3,12 @@ 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.Pi.Basic import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.GroupWithZero.Action.Basic import Mathlib.Algebra.GroupWithZero.Action.Units import Mathlib.Algebra.Module.Defs -import Mathlib.Algebra.Ring.Opposite import Mathlib.Algebra.NoZeroSMulDivisors.Defs +import Mathlib.Algebra.Ring.Opposite import Mathlib.Data.Set.Pairwise.Basic /-! @@ -44,9 +43,10 @@ section Mul variable [Mul α] {s t u : Set α} {a : α} -@[to_additive] -theorem op_smul_set_subset_mul : a ∈ t → op a • s ⊆ s * t := - image_subset_image2_left +@[to_additive] lemma smul_set_subset_mul : a ∈ s → a • t ⊆ s * t := image_subset_image2_right + +open scoped RightActions in +@[to_additive] lemma op_smul_set_subset_mul : a ∈ t → s <• a ⊆ s * t := image_subset_image2_left @[to_additive] theorem image_op_smul : (op '' s) • t = t * s := by @@ -65,6 +65,13 @@ theorem mul_subset_iff_left : s * t ⊆ u ↔ ∀ a ∈ s, a • t ⊆ u := theorem mul_subset_iff_right : s * t ⊆ u ↔ ∀ b ∈ t, op b • s ⊆ u := image2_subset_iff_right +@[to_additive] lemma pair_mul (a b : α) (s : Set α) : {a, b} * s = a • s ∪ b • s := by + rw [insert_eq, union_mul, singleton_mul, singleton_mul]; rfl + +open scoped RightActions +@[to_additive] lemma mul_pair (s : Set α) (a b : α) : s * {a, b} = s <• a ∪ s <• b := by + rw [insert_eq, mul_union, mul_singleton, mul_singleton]; rfl + end Mul variable {s s₁ s₂ : Set α} {t t₁ t₂ : Set β} {a : α} {b : β} @@ -291,7 +298,7 @@ end IsLeftCancelMul section Group -variable [Group α] [MulAction α β] {s t A B : Set β} {a : α} {x : β} +variable [Group α] [MulAction α β] {s t A B : Set β} {a b : α} {x : β} @[to_additive (attr := simp)] theorem smul_mem_smul_set_iff : a • x ∈ a • s ↔ x ∈ s := @@ -305,6 +312,10 @@ theorem mem_smul_set_iff_inv_smul_mem : x ∈ a • A ↔ a⁻¹ • x ∈ A := theorem mem_inv_smul_set_iff : x ∈ a⁻¹ • A ↔ a • x ∈ A := by simp only [← image_smul, mem_image, inv_smul_eq_iff, exists_eq_right] +@[to_additive (attr := simp)] +lemma mem_smul_set_inv {s : Set α} : a ∈ b • s⁻¹ ↔ b ∈ a • s := by + simp [mem_smul_set_iff_inv_smul_mem] + @[to_additive] theorem preimage_smul (a : α) (t : Set β) : (fun x ↦ a • x) ⁻¹' t = a⁻¹ • t := ((MulAction.toPerm a).symm.image_eq_preimage _).symm @@ -423,8 +434,35 @@ attribute [deprecated disjoint_smul_set (since := "2024-10-18")] smul_set_disjoi attribute [deprecated disjoint_vadd_set (since := "2024-10-18")] vadd_set_disjoint_iff +/-- Any intersection of translates of two sets `s` and `t` can be covered by a single translate of +`(s⁻¹ * s) ∩ (t⁻¹ * t)`. + +This is useful to show that the intersection of approximate subgroups is an approximate subgroup. -/ +@[to_additive +"Any intersection of translates of two sets `s` and `t` can be covered by a single translate of +`(-s + s) ∩ (-t + t)`. + +This is useful to show that the intersection of approximate subgroups is an approximate subgroup."] +lemma exists_smul_inter_smul_subset_smul_inv_mul_inter_inv_mul (s t : Set α) (a b : α) : + ∃ z : α, a • s ∩ b • t ⊆ z • ((s⁻¹ * s) ∩ (t⁻¹ * t)) := by + obtain hAB | ⟨z, hzA, hzB⟩ := (a • s ∩ b • t).eq_empty_or_nonempty + · exact ⟨1, by simp [hAB]⟩ + refine ⟨z, ?_⟩ + calc + a • s ∩ b • t ⊆ (z • s⁻¹) * s ∩ ((z • t⁻¹) * t) := by + gcongr <;> apply smul_set_subset_mul <;> simpa + _ = z • ((s⁻¹ * s) ∩ (t⁻¹ * t)) := by simp_rw [Set.smul_set_inter, smul_mul_assoc] + end Group +section Monoid +variable [Monoid α] [MulAction α β] {s : Set β} {a : α} {b : β} + +@[simp] lemma mem_invOf_smul_set [Invertible a] : b ∈ ⅟a • s ↔ a • b ∈ s := + mem_inv_smul_set_iff (a := unitOfInvertible a) + +end Monoid + section Group variable [Group α] [CommGroup β] [FunLike F α β] [MonoidHomClass F α β] From 19f864c3bad504cc6d55e314e85eab6a7f9e2ea3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Thu, 28 Nov 2024 14:24:39 +0000 Subject: [PATCH 383/829] feat(MeasureTheory): integrability lemmas (#19439) Mostly about the relation between `Integrable` and the Lebesgue integral for real/ennreal functions. From the TestingLowerBounds project. Co-authored-by: Lorenzo Luccioli --- Mathlib/MeasureTheory/Function/L1Space.lean | 95 ++++++++++++++++++--- 1 file changed, 82 insertions(+), 13 deletions(-) diff --git a/Mathlib/MeasureTheory/Function/L1Space.lean b/Mathlib/MeasureTheory/Function/L1Space.lean index b003cc9a9db03..722d303baab10 100644 --- a/Mathlib/MeasureTheory/Function/L1Space.lean +++ b/Mathlib/MeasureTheory/Function/L1Space.lean @@ -224,19 +224,30 @@ theorem hasFiniteIntegral_norm_iff (f : α → β) : HasFiniteIntegral (fun a => ‖f a‖) μ ↔ HasFiniteIntegral f μ := hasFiniteIntegral_congr' <| Eventually.of_forall fun x => norm_norm (f x) -theorem hasFiniteIntegral_toReal_of_lintegral_ne_top {f : α → ℝ≥0∞} (hf : (∫⁻ x, f x ∂μ) ≠ ∞) : - HasFiniteIntegral (fun x => (f x).toReal) μ := by - have : - ∀ x, (‖(f x).toReal‖₊ : ℝ≥0∞) = ENNReal.ofNNReal ⟨(f x).toReal, ENNReal.toReal_nonneg⟩ := by - intro x +theorem hasFiniteIntegral_toReal_of_lintegral_ne_top {f : α → ℝ≥0∞} (hf : ∫⁻ x, f x ∂μ ≠ ∞) : + HasFiniteIntegral (fun x ↦ (f x).toReal) μ := by + have h x : (‖(f x).toReal‖₊ : ℝ≥0∞) = ENNReal.ofNNReal ⟨(f x).toReal, ENNReal.toReal_nonneg⟩ := by rw [Real.nnnorm_of_nonneg] - simp_rw [HasFiniteIntegral, this] + simp_rw [HasFiniteIntegral, h] refine lt_of_le_of_lt (lintegral_mono fun x => ?_) (lt_top_iff_ne_top.2 hf) by_cases hfx : f x = ∞ · simp [hfx] · lift f x to ℝ≥0 using hfx with fx h simp [← h, ← NNReal.coe_le_coe] +lemma hasFiniteIntegral_toReal_iff {f : α → ℝ≥0∞} (hf_ne_top : ∀ᵐ x ∂μ, f x ≠ ∞) : + HasFiniteIntegral (fun x ↦ (f x).toReal) μ ↔ ∫⁻ x, f x ∂μ ≠ ∞ := by + have h_eq x : (‖(f x).toReal‖₊ : ℝ≥0∞) + = ENNReal.ofNNReal ⟨(f x).toReal, ENNReal.toReal_nonneg⟩ := by + rw [Real.nnnorm_of_nonneg] + simp_rw [HasFiniteIntegral, h_eq, lt_top_iff_ne_top] + convert Iff.rfl using 2 + refine lintegral_congr_ae ?_ + filter_upwards [hf_ne_top] with x hfx + lift f x to ℝ≥0 using hfx with fx h + simp only [coe_toReal, ENNReal.coe_inj] + rfl + theorem isFiniteMeasure_withDensity_ofReal {f : α → ℝ} (hfi : HasFiniteIntegral f μ) : IsFiniteMeasure (μ.withDensity fun x => ENNReal.ofReal <| f x) := by refine isFiniteMeasure_withDensity ((lintegral_mono fun x => ?_).trans_lt hfi).ne @@ -605,24 +616,53 @@ theorem integrable_finset_sum {ι} (s : Finset ι) {f : ι → α → β} (hf : ∀ i ∈ s, Integrable (f i) μ) : Integrable (fun a => ∑ i ∈ s, f i a) μ := by simpa only [← Finset.sum_apply] using integrable_finset_sum' s hf +/-- If `f` is integrable, then so is `-f`. +See `Integrable.neg'` for the same statement, but formulated with `x ↦ - f x` instead of `-f`. -/ theorem Integrable.neg {f : α → β} (hf : Integrable f μ) : Integrable (-f) μ := ⟨hf.aestronglyMeasurable.neg, hf.hasFiniteIntegral.neg⟩ +/-- If `f` is integrable, then so is `fun x ↦ - f x`. +See `Integrable.neg` for the same statement, but formulated with `-f` instead of `fun x ↦ - f x`. -/ +theorem Integrable.neg' {f : α → β} (hf : Integrable f μ) : Integrable (fun x ↦ - f x) μ := + ⟨hf.aestronglyMeasurable.neg, hf.hasFiniteIntegral.neg⟩ + @[simp] theorem integrable_neg_iff {f : α → β} : Integrable (-f) μ ↔ Integrable f μ := ⟨fun h => neg_neg f ▸ h.neg, Integrable.neg⟩ +/-- if `f` is integrable, then `f + g` is integrable iff `g` is. +See `integrable_add_iff_integrable_right'` for the same statement with `fun x ↦ f x + g x` instead +of `f + g`. -/ @[simp] lemma integrable_add_iff_integrable_right {f g : α → β} (hf : Integrable f μ) : Integrable (f + g) μ ↔ Integrable g μ := ⟨fun h ↦ show g = f + g + (-f) by simp only [add_neg_cancel_comm] ▸ h.add hf.neg, fun h ↦ hf.add h⟩ +/-- if `f` is integrable, then `fun x ↦ f x + g x` is integrable iff `g` is. +See `integrable_add_iff_integrable_right` for the same statement with `f + g` instead +of `fun x ↦ f x + g x`. -/ +@[simp] +lemma integrable_add_iff_integrable_right' {f g : α → β} (hf : Integrable f μ) : + Integrable (fun x ↦ f x + g x) μ ↔ Integrable g μ := + integrable_add_iff_integrable_right hf + +/-- if `f` is integrable, then `g + f` is integrable iff `g` is. +See `integrable_add_iff_integrable_left'` for the same statement with `fun x ↦ g x + f x` instead +of `g + f`. -/ @[simp] lemma integrable_add_iff_integrable_left {f g : α → β} (hf : Integrable f μ) : Integrable (g + f) μ ↔ Integrable g μ := by rw [add_comm, integrable_add_iff_integrable_right hf] +/-- if `f` is integrable, then `fun x ↦ g x + f x` is integrable iff `g` is. +See `integrable_add_iff_integrable_left'` for the same statement with `g + f` instead +of `fun x ↦ g x + f x`. -/ +@[simp] +lemma integrable_add_iff_integrable_left' {f g : α → β} (hf : Integrable f μ) : + Integrable (fun x ↦ g x + f x) μ ↔ Integrable g μ := + integrable_add_iff_integrable_left hf + lemma integrable_left_of_integrable_add_of_nonneg {f g : α → ℝ} (h_meas : AEStronglyMeasurable f μ) (hf : 0 ≤ᵐ[μ] f) (hg : 0 ≤ᵐ[μ] g) (h_int : Integrable (f + g) μ) : Integrable f μ := by @@ -650,12 +690,10 @@ lemma integrable_add_iff_of_nonpos {f g : α → ℝ} (h_meas : AEStronglyMeasur exact integrable_add_iff_of_nonneg h_meas.neg (hf.mono (fun _ ↦ neg_nonneg_of_nonpos)) (hg.mono (fun _ ↦ neg_nonneg_of_nonpos)) -@[simp] lemma integrable_add_const_iff [IsFiniteMeasure μ] {f : α → β} {c : β} : Integrable (fun x ↦ f x + c) μ ↔ Integrable f μ := integrable_add_iff_integrable_left (integrable_const _) -@[simp] lemma integrable_const_add_iff [IsFiniteMeasure μ] {f : α → β} {c : β} : Integrable (fun x ↦ c + f x) μ ↔ Integrable f μ := integrable_add_iff_integrable_right (integrable_const _) @@ -745,6 +783,21 @@ theorem integrable_of_norm_sub_le {f₀ f₁ : α → β} {g : α → ℝ} (hf _ ≤ ‖f₀ a‖ + g a := add_le_add_left ha _ Integrable.mono' (hf₀_i.norm.add hg_i) hf₁_m this +lemma integrable_of_le_of_le {f g₁ g₂ : α → ℝ} (hf : AEStronglyMeasurable f μ) + (h_le₁ : g₁ ≤ᵐ[μ] f) (h_le₂ : f ≤ᵐ[μ] g₂) + (h_int₁ : Integrable g₁ μ) (h_int₂ : Integrable g₂ μ) : + Integrable f μ := by + have : ∀ᵐ x ∂μ, ‖f x‖ ≤ max ‖g₁ x‖ ‖g₂ x‖ := by + filter_upwards [h_le₁, h_le₂] with x hx1 hx2 + simp only [Real.norm_eq_abs] + exact abs_le_max_abs_abs hx1 hx2 + have h_le_add : ∀ᵐ x ∂μ, ‖f x‖ ≤ ‖‖g₁ x‖ + ‖g₂ x‖‖ := by + filter_upwards [this] with x hx + refine hx.trans ?_ + conv_rhs => rw [Real.norm_of_nonneg (by positivity)] + exact max_le_add_of_nonneg (norm_nonneg _) (norm_nonneg _) + exact Integrable.mono (h_int₁.norm.add h_int₂.norm) hf h_le_add + theorem Integrable.prod_mk {f : α → β} {g : α → γ} (hf : Integrable f μ) (hg : Integrable g μ) : Integrable (fun x => (f x, g x)) μ := ⟨hf.aestronglyMeasurable.prod_mk hg.aestronglyMeasurable, @@ -1003,17 +1056,31 @@ theorem withDensitySMulLI_apply {f : α → ℝ≥0} (f_meas : Measurable f) end +section ENNReal + theorem mem_ℒ1_toReal_of_lintegral_ne_top {f : α → ℝ≥0∞} (hfm : AEMeasurable f μ) - (hfi : (∫⁻ x, f x ∂μ) ≠ ∞) : Memℒp (fun x => (f x).toReal) 1 μ := by + (hfi : ∫⁻ x, f x ∂μ ≠ ∞) : Memℒp (fun x ↦ (f x).toReal) 1 μ := by rw [Memℒp, eLpNorm_one_eq_lintegral_nnnorm] - exact - ⟨(AEMeasurable.ennreal_toReal hfm).aestronglyMeasurable, - hasFiniteIntegral_toReal_of_lintegral_ne_top hfi⟩ + exact ⟨(AEMeasurable.ennreal_toReal hfm).aestronglyMeasurable, + hasFiniteIntegral_toReal_of_lintegral_ne_top hfi⟩ theorem integrable_toReal_of_lintegral_ne_top {f : α → ℝ≥0∞} (hfm : AEMeasurable f μ) - (hfi : (∫⁻ x, f x ∂μ) ≠ ∞) : Integrable (fun x => (f x).toReal) μ := + (hfi : ∫⁻ x, f x ∂μ ≠ ∞) : Integrable (fun x ↦ (f x).toReal) μ := memℒp_one_iff_integrable.1 <| mem_ℒ1_toReal_of_lintegral_ne_top hfm hfi +lemma integrable_toReal_iff {f : α → ℝ≥0∞} (hf : AEMeasurable f μ) (hf_ne_top : ∀ᵐ x ∂μ, f x ≠ ∞) : + Integrable (fun x ↦ (f x).toReal) μ ↔ ∫⁻ x, f x ∂μ ≠ ∞ := by + rw [Integrable, hasFiniteIntegral_toReal_iff hf_ne_top] + simp only [hf.ennreal_toReal.aestronglyMeasurable, ne_eq, true_and] + +lemma lintegral_ofReal_ne_top_iff_integrable {f : α → ℝ} + (hfm : AEStronglyMeasurable f μ) (hf : 0 ≤ᵐ[μ] f) : + ∫⁻ a, ENNReal.ofReal (f a) ∂μ ≠ ∞ ↔ Integrable f μ := by + rw [Integrable, hasFiniteIntegral_iff_ofReal hf] + simp [hfm] + +end ENNReal + section PosPart /-! ### Lemmas used for defining the positive part of an `L¹` function -/ @@ -1456,3 +1523,5 @@ lemma integrable_prod {f : α → E × F} : ⟨fun h ↦ ⟨h.fst, h.snd⟩, fun h ↦ h.1.prod_mk h.2⟩ end MeasureTheory + +set_option linter.style.longFile 1700 From 0fc63c9122ae798d30bdba64a7e7bf4f960018e1 Mon Sep 17 00:00:00 2001 From: Herman Chau Date: Thu, 28 Nov 2024 14:24:41 +0000 Subject: [PATCH 384/829] doc: Fix minor typos in Set.Operations and Set.Basic (#19549) --- Mathlib/Data/Set/Finite/Basic.lean | 4 ++-- Mathlib/Data/Set/Operations.lean | 16 ++++++++-------- 2 files changed, 10 insertions(+), 10 deletions(-) diff --git a/Mathlib/Data/Set/Finite/Basic.lean b/Mathlib/Data/Set/Finite/Basic.lean index 800c3708d2d18..0ee4e5ce237b5 100644 --- a/Mathlib/Data/Set/Finite/Basic.lean +++ b/Mathlib/Data/Set/Finite/Basic.lean @@ -324,8 +324,8 @@ instance (priority := 100) fintypeInsert' (a : α) (s : Set α) [Decidable <| a instance fintypeImage [DecidableEq β] (s : Set α) (f : α → β) [Fintype s] : Fintype (f '' s) := Fintype.ofFinset (s.toFinset.image f) <| by simp -/-- If a function `f` has a partial inverse and sends a set `s` to a set with `[Fintype]` instance, -then `s` has a `Fintype` structure as well. -/ +/-- If a function `f` has a partial inverse `g` and the image of `s` under `f` is a set with +a `Fintype` instance, then `s` has a `Fintype` structure as well. -/ def fintypeOfFintypeImage (s : Set α) {f : α → β} {g} (I : IsPartialInv f g) [Fintype (f '' s)] : Fintype s := Fintype.ofFinset ⟨_, (f '' s).toFinset.2.filterMap g <| injective_of_isPartialInv_right I⟩ diff --git a/Mathlib/Data/Set/Operations.lean b/Mathlib/Data/Set/Operations.lean index bed6dc9156d8b..8abbfac04aea9 100644 --- a/Mathlib/Data/Set/Operations.lean +++ b/Mathlib/Data/Set/Operations.lean @@ -213,8 +213,8 @@ section Pi variable {ι : Type*} {α : ι → Type*} /-- Given an index set `ι` and a family of sets `t : Π i, Set (α i)`, `pi s t` -is the set of dependent functions `f : Πa, π a` such that `f a` belongs to `t a` -whenever `a ∈ s`. -/ +is the set of dependent functions `f : Πa, π a` such that `f i` belongs to `t i` +whenever `i ∈ s`. -/ def pi (s : Set ι) (t : ∀ i, Set (α i)) : Set (∀ i, α i) := {f | ∀ i ∈ s, f i ∈ t i} variable {s : Set ι} {t : ∀ i, Set (α i)} {f : ∀ i, α i} @@ -228,7 +228,7 @@ end Pi /-- Two functions `f₁ f₂ : α → β` are equal on `s` if `f₁ x = f₂ x` for all `x ∈ s`. -/ def EqOn (f₁ f₂ : α → β) (s : Set α) : Prop := ∀ ⦃x⦄, x ∈ s → f₁ x = f₂ x -/-- `MapsTo f a b` means that the image of `a` is contained in `b`. -/ +/-- `MapsTo f s t` means that the image of `s` is contained in `t`. -/ def MapsTo (f : α → β) (s : Set α) (t : Set β) : Prop := ∀ ⦃x⦄, x ∈ s → f x ∈ t theorem mapsTo_image (f : α → β) (s : Set α) : MapsTo f s (f '' s) := fun _ ↦ mem_image_of_mem f @@ -245,26 +245,26 @@ def MapsTo.restrict (f : α → β) (s : Set α) (t : Set β) (h : MapsTo f s t) def restrictPreimage (t : Set β) (f : α → β) : f ⁻¹' t → t := (Set.mapsTo_preimage f t).restrict _ _ _ -/-- `f` is injective on `a` if the restriction of `f` to `a` is injective. -/ +/-- `f` is injective on `s` if the restriction of `f` to `s` is injective. -/ def InjOn (f : α → β) (s : Set α) : Prop := ∀ ⦃x₁ : α⦄, x₁ ∈ s → ∀ ⦃x₂ : α⦄, x₂ ∈ s → f x₁ = f x₂ → x₁ = x₂ /-- The graph of a function `f : α → β` on a set `s`. -/ def graphOn (f : α → β) (s : Set α) : Set (α × β) := (fun x ↦ (x, f x)) '' s -/-- `f` is surjective from `a` to `b` if `b` is contained in the image of `a`. -/ +/-- `f` is surjective from `s` to `t` if `t` is contained in the image of `s`. -/ def SurjOn (f : α → β) (s : Set α) (t : Set β) : Prop := t ⊆ f '' s /-- `f` is bijective from `s` to `t` if `f` is injective on `s` and `f '' s = t`. -/ def BijOn (f : α → β) (s : Set α) (t : Set β) : Prop := MapsTo f s t ∧ InjOn f s ∧ SurjOn f s t -/-- `g` is a left inverse to `f` on `a` means that `g (f x) = x` for all `x ∈ a`. -/ +/-- `g` is a left inverse to `f` on `s` means that `g (f x) = x` for all `x ∈ s`. -/ def LeftInvOn (f' : β → α) (f : α → β) (s : Set α) : Prop := ∀ ⦃x⦄, x ∈ s → f' (f x) = x -/-- `g` is a right inverse to `f` on `b` if `f (g x) = x` for all `x ∈ b`. -/ +/-- `g` is a right inverse to `f` on `t` if `f (g x) = x` for all `x ∈ t`. -/ abbrev RightInvOn (f' : β → α) (f : α → β) (t : Set β) : Prop := LeftInvOn f f' t -/-- `g` is an inverse to `f` viewed as a map from `a` to `b` -/ +/-- `g` is an inverse to `f` viewed as a map from `s` to `t` -/ def InvOn (g : β → α) (f : α → β) (s : Set α) (t : Set β) : Prop := LeftInvOn g f s ∧ RightInvOn g f t From 0823ec54412c06bc182ad1dbca1c23205cbd8ff5 Mon Sep 17 00:00:00 2001 From: Arend Mellendijk Date: Thu, 28 Nov 2024 15:29:10 +0000 Subject: [PATCH 385/829] feat: Define `Nat.finMulAntidiagonal`, the `Finset` of tuples of naturals with a fixed product. (#10668) The result `card_finMulAntidiagonal_of_squarefree` is central to controlling the error term in my formalisation of the Selberg sieve. (specifically for lemma 3.1 in [Heath-Brown's notes](https://arxiv.org/abs/math/0209360)) The definition is adapted from [Finset.finAntidiagonal](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Data/Finset/PiAntidiagonal.html#Finset.finAntidiagonal), though it is held back a bit by `Nat.divisorsAntidiagonal` which has a time complexity of $O(n^2)$. This definition is limited to `Nat` because it's not clear how to handle 0 in higher generality. ([zulip](https://leanprover.zulipchat.com/#narrow/stream/144837-PR-reviews/topic/.237486.20antidiagonal/near/395886479)) Co-authored-by: Arend Mellendijk Co-authored-by: Eric Wieser --- Mathlib.lean | 1 + Mathlib/Algebra/Order/Antidiag/Nat.lean | 242 ++++++++++++++++++ .../Order/BigOperators/Ring/Finset.lean | 5 + Mathlib/Data/PNat/Defs.lean | 2 +- Mathlib/Logic/Equiv/Basic.lean | 1 + 5 files changed, 250 insertions(+), 1 deletion(-) create mode 100644 Mathlib/Algebra/Order/Antidiag/Nat.lean diff --git a/Mathlib.lean b/Mathlib.lean index 4ace0a7b6e35c..07114d775aeaa 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -615,6 +615,7 @@ import Mathlib.Algebra.Order.AddGroupWithTop import Mathlib.Algebra.Order.AddTorsor import Mathlib.Algebra.Order.Algebra import Mathlib.Algebra.Order.Antidiag.Finsupp +import Mathlib.Algebra.Order.Antidiag.Nat import Mathlib.Algebra.Order.Antidiag.Pi import Mathlib.Algebra.Order.Antidiag.Prod import Mathlib.Algebra.Order.Archimedean.Basic diff --git a/Mathlib/Algebra/Order/Antidiag/Nat.lean b/Mathlib/Algebra/Order/Antidiag/Nat.lean new file mode 100644 index 0000000000000..40c8db2fe9e60 --- /dev/null +++ b/Mathlib/Algebra/Order/Antidiag/Nat.lean @@ -0,0 +1,242 @@ +/- +Copyright (c) 2024 Arend Mellendijk. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Arend Mellendijk +-/ + +import Mathlib.Algebra.Order.Antidiag.Pi +import Mathlib.NumberTheory.ArithmeticFunction + +/-! +# Sets of tuples with a fixed product + +This file defines the finite set of `d`-tuples of natural numbers with a fixed product `n` as +`Nat.finMulAntidiagonal`. + +## Main Results +* There are `d^(ω n)` ways to write `n` as a product of `d` natural numbers, when `n` is squarefree +(`card_finMulAntidiagonal_of_squarefree`) +-/ + +open Finset +open scoped BigOperators ArithmeticFunction +namespace PNat + +instance instHasAntidiagonal : Finset.HasAntidiagonal (Additive ℕ+) := + /- The set of divisors of a positive natural number. +This is `Nat.divisorsAntidiagonal` without a special case for `n = 0`. -/ + let divisorsAntidiagonal (n : ℕ+) : Finset (ℕ+ × ℕ+) := + (Nat.divisorsAntidiagonal n).attach.map + ⟨fun x => + (⟨x.val.1, Nat.pos_of_mem_divisors <| Nat.fst_mem_divisors_of_mem_antidiagonal x.prop⟩, + ⟨x.val.2, Nat.pos_of_mem_divisors <| Nat.snd_mem_divisors_of_mem_antidiagonal x.prop⟩), + fun _ _ h => Subtype.ext <| Prod.ext (congr_arg (·.1.val) h) (congr_arg (·.2.val) h)⟩ + + have mem_divisorsAntidiagonal {n : ℕ+} (x : ℕ+ × ℕ+) : + x ∈ divisorsAntidiagonal n ↔ x.1 * x.2 = n := by + simp_rw [divisorsAntidiagonal, Finset.mem_map, Finset.mem_attach, Function.Embedding.coeFn_mk, + Prod.ext_iff, true_and, ← coe_inj, Subtype.exists] + aesop + { antidiagonal := fun n ↦ divisorsAntidiagonal (Additive.toMul n) |>.map + (.prodMap (Additive.ofMul.toEmbedding) (Additive.ofMul.toEmbedding)) + mem_antidiagonal := by simp [← ofMul_mul, mem_divisorsAntidiagonal] } + +end PNat + +namespace Nat + +/-- The `Finset` of all `d`-tuples of natural numbers whose product is `n`. Defined to be `∅` when +`n = 0`. -/ +def finMulAntidiag (d : ℕ) (n : ℕ) : Finset (Fin d → ℕ) := + if hn : 0 < n then + (Finset.finAntidiagonal d (Additive.ofMul (α := ℕ+) ⟨n, hn⟩)).map <| + .arrowCongrRight <| Additive.toMul.toEmbedding.trans <| ⟨PNat.val, PNat.coe_injective⟩ + else + ∅ + +@[simp] +theorem mem_finMulAntidiag {d n : ℕ} {f : Fin d → ℕ} : + f ∈ finMulAntidiag d n ↔ ∏ i, f i = n ∧ n ≠ 0 := by + unfold finMulAntidiag + split_ifs with h + · simp_rw [mem_map, mem_finAntidiagonal, Function.Embedding.arrowCongrRight_apply, + Function.comp_def, Function.Embedding.trans_apply, Equiv.coe_toEmbedding, + Function.Embedding.coeFn_mk, ← Additive.ofMul.symm_apply_eq, Additive.ofMul_symm_eq, + toMul_sum, (Equiv.piCongrRight fun _=> Additive.ofMul).surjective.exists, + Equiv.piCongrRight_apply, Pi.map_apply, toMul_ofMul, ← PNat.coe_inj, PNat.mk_coe, + PNat.coe_prod] + constructor + · rintro ⟨a, ha_mem, rfl⟩ + exact ⟨ha_mem, h.ne.symm⟩ + · rintro ⟨rfl, _⟩ + refine ⟨fun i ↦ ⟨f i, ?_⟩, rfl, funext fun _ => rfl⟩ + apply Nat.pos_of_ne_zero + exact Finset.prod_ne_zero_iff.mp h.ne.symm _ (mem_univ _) + · simp only [not_lt, nonpos_iff_eq_zero] at h + simp only [h, not_mem_empty, ne_eq, not_true_eq_false, and_false] + +@[simp] +theorem finMulAntidiag_zero_right (d : ℕ) : + finMulAntidiag d 0 = ∅ := rfl + +theorem finMulAntidiag_one {d : ℕ} : + finMulAntidiag d 1 = {fun _ => 1} := by + ext f + simp only [mem_finMulAntidiag, and_true, mem_singleton] + constructor + · intro ⟨hf, _⟩; ext i + rw [← Nat.dvd_one, ← hf] + exact dvd_prod_of_mem f (mem_univ _) + · rintro rfl + simp only [prod_const_one, implies_true, ne_eq, one_ne_zero, not_false_eq_true, + and_self] + +theorem finMulAntidiag_zero_left {n : ℕ} (hn : n ≠ 1) : + finMulAntidiag 0 n = ∅ := by + ext + simp [hn.symm] + +theorem dvd_of_mem_finMulAntidiag {n d : ℕ} {f : Fin d → ℕ} (hf : f ∈ finMulAntidiag d n) + (i : Fin d) : f i ∣ n := by + rw [mem_finMulAntidiag] at hf + rw [← hf.1] + exact dvd_prod_of_mem f (mem_univ i) + +theorem ne_zero_of_mem_finMulAntidiag {d n : ℕ} {f : Fin d → ℕ} + (hf : f ∈ finMulAntidiag d n) (i : Fin d) : f i ≠ 0 := + ne_zero_of_dvd_ne_zero (mem_finMulAntidiag.mp hf).2 (dvd_of_mem_finMulAntidiag hf i) + +theorem prod_eq_of_mem_finMulAntidiag {d n : ℕ} {f : Fin d → ℕ} + (hf : f ∈ finMulAntidiag d n) : ∏ i, f i = n := + (mem_finMulAntidiag.mp hf).1 + +theorem finMulAntidiag_eq_piFinset_divisors_filter {d m n : ℕ} (hmn : m ∣ n) (hn : n ≠ 0) : + finMulAntidiag d m = + {f ∈ Fintype.piFinset fun _ : Fin d => n.divisors | ∏ i, f i = m} := by + ext f + simp only [mem_univ, not_true, IsEmpty.forall_iff, implies_true, ne_eq, true_and, + Fintype.mem_piFinset, mem_divisors, Nat.isUnit_iff, mem_filter] + constructor + · intro hf + refine ⟨?_, prod_eq_of_mem_finMulAntidiag hf⟩ + exact fun i => ⟨(dvd_of_mem_finMulAntidiag hf i).trans hmn, hn⟩ + · rw [mem_finMulAntidiag] + exact fun ⟨_, hprod⟩ => ⟨hprod, ne_zero_of_dvd_ne_zero hn hmn⟩ + +lemma image_apply_finMulAntidiag {d n : ℕ} {i : Fin d} (hd : d ≠ 1) : + (finMulAntidiag d n).image (fun f => f i) = divisors n := by + ext k + simp only [mem_image, ne_eq, mem_divisors, Nat.isUnit_iff] + constructor + · rintro ⟨f, hf, rfl⟩ + exact ⟨dvd_of_mem_finMulAntidiag hf _, (mem_finMulAntidiag.mp hf).2⟩ + · simp_rw [mem_finMulAntidiag] + rintro ⟨⟨r, rfl⟩, hn⟩ + have hs : Nontrivial (Fin d) := by + rw [Fin.nontrivial_iff_two_le] + obtain rfl | hd' := eq_or_ne d 0 + · exact i.elim0 + omega + obtain ⟨i', hi_ne⟩ := exists_ne i + use fun j => if j = i then k else if j = i' then r else 1 + simp only [ite_true, and_true, hn] + rw [← Finset.mul_prod_erase (a:=i) (h:=mem_univ _), + ← Finset.mul_prod_erase (a:= i')] + · rw [if_neg hi_ne, if_pos rfl, if_pos rfl, prod_eq_one] + · refine ⟨by ring, hn⟩ + intro j hj + simp only [mem_erase, ne_eq, mem_univ, and_true] at hj + rw [if_neg hj.1, if_neg hj.2] + exact mem_erase.mpr ⟨hi_ne, mem_univ _⟩ + +lemma image_piFinTwoEquiv_finMulAntidiag {n : ℕ} : + (finMulAntidiag 2 n).image (piFinTwoEquiv <| fun _ => ℕ) = divisorsAntidiagonal n := by + ext x + simp [(piFinTwoEquiv <| fun _ => ℕ).symm.surjective.exists] + +lemma finMulAntidiag_exists_unique_prime_dvd {d n p : ℕ} (hn : Squarefree n) + (hp : p ∈ n.primeFactorsList) (f : Fin d → ℕ) (hf : f ∈ finMulAntidiag d n) : + ∃! i, p ∣ f i := by + rw [mem_finMulAntidiag] at hf + rw [mem_primeFactorsList hf.2, ← hf.1, hp.1.prime.dvd_finset_prod_iff] at hp + obtain ⟨i, his, hi⟩ := hp.2 + refine ⟨i, hi, ?_⟩ + intro j hj + by_contra hij + apply Nat.Prime.not_coprime_iff_dvd.mpr ⟨p, hp.1, hi, hj⟩ + apply Nat.coprime_of_squarefree_mul + apply hn.squarefree_of_dvd + rw [← hf.1, ← Finset.mul_prod_erase _ _ (his), + ← Finset.mul_prod_erase _ _ (mem_erase.mpr ⟨hij, mem_univ _⟩), ← mul_assoc] + apply Nat.dvd_mul_right + +private def primeFactorsPiBij (d n : ℕ) : + ∀ f ∈ (n.primeFactors.pi fun _ => (univ : Finset <| Fin d)), Fin d → ℕ := + fun f _ i => ∏ p ∈ {p ∈ n.primeFactors.attach | f p.1 p.2 = i} , p + +private theorem primeFactorsPiBij_img (d n : ℕ) (hn : Squarefree n) + (f : (p : ℕ) → p ∈ n.primeFactors → Fin d) (hf : f ∈ pi n.primeFactors fun _ => univ) : + Nat.primeFactorsPiBij d n f hf ∈ finMulAntidiag d n := by + rw [mem_finMulAntidiag] + refine ⟨?_, hn.ne_zero⟩ + unfold Nat.primeFactorsPiBij + rw [prod_fiberwise_of_maps_to, prod_attach (f := fun x => x)] + · apply prod_primeFactors_of_squarefree hn + · apply fun _ _ => mem_univ _ + +private theorem primeFactorsPiBij_inj (d n : ℕ) + (f : (p : ℕ) → p ∈ n.primeFactors → Fin d) (hf : f ∈ pi n.primeFactors fun _ => univ) + (g : (p : ℕ) → p ∈ n.primeFactors → Fin d) (hg : g ∈ pi n.primeFactors fun _ => univ) : + Nat.primeFactorsPiBij d n f hf = Nat.primeFactorsPiBij d n g hg → f = g := by + contrapose! + simp_rw [Function.ne_iff] + intro ⟨p, hp, hfg⟩ + use f p hp + dsimp only [Nat.primeFactorsPiBij] + apply ne_of_mem_of_not_mem (s := {x | p ∣ x}) <;> simp_rw [Set.mem_setOf_eq] + · rw [Finset.prod_filter] + convert Finset.dvd_prod_of_mem _ (mem_attach (n.primeFactors) ⟨p, hp⟩) + rw [if_pos rfl] + · rw [mem_primeFactors] at hp + rw [Prime.dvd_finset_prod_iff hp.1.prime] + push_neg + intro q hq + rw [Nat.prime_dvd_prime_iff_eq hp.1 (Nat.prime_of_mem_primeFactorsList + <| List.mem_toFinset.mp q.2)] + rintro rfl + rw [(mem_filter.mp hq).2] at hfg + exact hfg rfl + +private theorem primeFactorsPiBij_surj (d n : ℕ) (hn : Squarefree n) + (t : Fin d → ℕ) (ht : t ∈ finMulAntidiag d n) : ∃ (g : _) + (hg : g ∈ pi n.primeFactors fun _ => univ), Nat.primeFactorsPiBij d n g hg = t := by + have exists_unique := fun (p : ℕ) (hp : p ∈ n.primeFactors) => + (finMulAntidiag_exists_unique_prime_dvd hn + (mem_primeFactors_iff_mem_primeFactorsList.mp hp) t ht) + choose f hf hf_unique using exists_unique + refine ⟨f, ?_, ?_⟩ + · simp only [mem_pi, mem_univ, forall_true_iff] + funext i + have : t i ∣ n := dvd_of_mem_finMulAntidiag ht _ + trans (∏ p ∈ n.primeFactors.attach, if p.1 ∣ t i then p else 1) + · rw [Nat.primeFactorsPiBij, ← prod_filter] + congr + ext ⟨p, hp⟩ + refine ⟨by rintro rfl; apply hf, fun h => (hf_unique p hp i h).symm⟩ + rw [prod_attach (f:=fun p => if p ∣ t i then p else 1), ← Finset.prod_filter] + rw [primeFactors_filter_dvd_of_dvd hn.ne_zero this] + exact prod_primeFactors_of_squarefree <| hn.squarefree_of_dvd this + +private theorem card_finMulAntidiag_pi (d n : ℕ) (hn : Squarefree n) : + #(n.primeFactors.pi fun _ => (univ : Finset <| Fin d)) = + #(finMulAntidiag d n) := by + apply Finset.card_bij (Nat.primeFactorsPiBij d n) (primeFactorsPiBij_img d n hn) + (primeFactorsPiBij_inj d n) (primeFactorsPiBij_surj d n hn) + +theorem card_finMulAntidiag_of_squarefree {d n : ℕ} (hn : Squarefree n) : + #(finMulAntidiag d n) = d ^ ω n := by + rw [← card_finMulAntidiag_pi d n hn, Finset.card_pi, Finset.prod_const, + ArithmeticFunction.cardDistinctFactors_apply, ← List.card_toFinset, toFinset_factors, + Finset.card_fin] + +end Nat diff --git a/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean b/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean index 1c522e8948d04..6a0c16f4bea90 100644 --- a/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean @@ -132,6 +132,11 @@ lemma abs_prod [LinearOrderedCommRing R] (s : Finset ι) (f : ι → R) : |∏ x ∈ s, f x| = ∏ x ∈ s, |f x| := map_prod absHom _ _ +@[simp, norm_cast] +theorem PNat.coe_prod {ι : Type*} (f : ι → ℕ+) (s : Finset ι) : + ↑(∏ i ∈ s, f i) = (∏ i ∈ s, f i : ℕ) := + map_prod PNat.coeMonoidHom _ _ + section CanonicallyOrderedCommSemiring variable [CanonicallyOrderedCommSemiring R] {f g h : ι → R} {s : Finset ι} {i : ι} diff --git a/Mathlib/Data/PNat/Defs.lean b/Mathlib/Data/PNat/Defs.lean index 6eb644184ac56..afc7437cfeddc 100644 --- a/Mathlib/Data/PNat/Defs.lean +++ b/Mathlib/Data/PNat/Defs.lean @@ -114,7 +114,7 @@ theorem pos (n : ℕ+) : 0 < (n : ℕ) := theorem eq {m n : ℕ+} : (m : ℕ) = n → m = n := Subtype.eq -theorem coe_injective : Function.Injective (fun (a : ℕ+) => (a : ℕ)) := +theorem coe_injective : Function.Injective PNat.val := Subtype.coe_injective @[simp] diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 0d001f766b84f..070edf6a55b4e 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -627,6 +627,7 @@ section /-- A family of equivalences `∀ a, β₁ a ≃ β₂ a` generates an equivalence between `∀ a, β₁ a` and `∀ a, β₂ a`. -/ +@[simps] def piCongrRight {β₁ β₂ : α → Sort*} (F : ∀ a, β₁ a ≃ β₂ a) : (∀ a, β₁ a) ≃ (∀ a, β₂ a) := ⟨Pi.map fun a ↦ F a, Pi.map fun a ↦ (F a).symm, fun H => funext <| by simp, fun H => funext <| by simp⟩ From 9c17b7f962387696e6c8dbe3cc6a78249b6306cd Mon Sep 17 00:00:00 2001 From: sgouezel Date: Thu, 28 Nov 2024 15:29:11 +0000 Subject: [PATCH 386/829] feat: manifold derivative of operations in vector spaces (#18849) We have a comprehensive API for higher smoothness in the manifold sense of functions on vector spaces. We copy this API to deal with manifold differentiability. --- Mathlib.lean | 1 + .../Manifold/MFDeriv/NormedSpace.lean | 277 ++++++++++++++++++ 2 files changed, 278 insertions(+) create mode 100644 Mathlib/Geometry/Manifold/MFDeriv/NormedSpace.lean diff --git a/Mathlib.lean b/Mathlib.lean index 07114d775aeaa..d1faa91f26155 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3063,6 +3063,7 @@ import Mathlib.Geometry.Manifold.MFDeriv.Atlas import Mathlib.Geometry.Manifold.MFDeriv.Basic import Mathlib.Geometry.Manifold.MFDeriv.Defs import Mathlib.Geometry.Manifold.MFDeriv.FDeriv +import Mathlib.Geometry.Manifold.MFDeriv.NormedSpace import Mathlib.Geometry.Manifold.MFDeriv.SpecificFunctions import Mathlib.Geometry.Manifold.MFDeriv.Tangent import Mathlib.Geometry.Manifold.MFDeriv.UniqueDifferential diff --git a/Mathlib/Geometry/Manifold/MFDeriv/NormedSpace.lean b/Mathlib/Geometry/Manifold/MFDeriv/NormedSpace.lean new file mode 100644 index 0000000000000..7b650ffa5716c --- /dev/null +++ b/Mathlib/Geometry/Manifold/MFDeriv/NormedSpace.lean @@ -0,0 +1,277 @@ +/- +Copyright (c) 2024 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel, Floris van Doorn +-/ +import Mathlib.Geometry.Manifold.ContMDiff.NormedSpace +import Mathlib.Geometry.Manifold.MFDeriv.FDeriv + +/-! ## Equivalence of manifold differentiability with the basic definition for functions between +vector spaces + +The API in this file is mostly copied from `Mathlib.Geometry.Manifold.ContMDiff.NormedSpace`, +providing the same statements for higher smoothness. In this file, we do the same for +differentiability. + +-/ + +open Set ChartedSpace SmoothManifoldWithCorners +open scoped Topology Manifold + +variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] + -- declare a charted space `M` over the pair `(E, H)`. + {E : Type*} + [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] + {I : ModelWithCorners 𝕜 E H} {M : Type*} [TopologicalSpace M] [ChartedSpace H M] + -- declare a smooth manifold `M'` over the pair `(E', H')`. + {E' : Type*} + [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] + {I' : ModelWithCorners 𝕜 E' H'} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] + [SmoothManifoldWithCorners I' M'] + -- declare a smooth manifold `N` over the pair `(F, G)`. + {F : Type*} + [NormedAddCommGroup F] [NormedSpace 𝕜 F] {G : Type*} [TopologicalSpace G] + {J : ModelWithCorners 𝕜 F G} {N : Type*} [TopologicalSpace N] [ChartedSpace G N] + [SmoothManifoldWithCorners J N] + -- declare a smooth manifold `N'` over the pair `(F', G')`. + {F' : Type*} + [NormedAddCommGroup F'] [NormedSpace 𝕜 F'] {G' : Type*} [TopologicalSpace G'] + {J' : ModelWithCorners 𝕜 F' G'} {N' : Type*} [TopologicalSpace N'] [ChartedSpace G' N'] + [SmoothManifoldWithCorners J' N'] + -- F₁, F₂, F₃, F₄ are normed spaces + {F₁ : Type*} [NormedAddCommGroup F₁] [NormedSpace 𝕜 F₁] {F₂ : Type*} [NormedAddCommGroup F₂] + [NormedSpace 𝕜 F₂] {F₃ : Type*} [NormedAddCommGroup F₃] [NormedSpace 𝕜 F₃] {F₄ : Type*} + [NormedAddCommGroup F₄] [NormedSpace 𝕜 F₄] + -- declare functions, sets, points and smoothness indices + {f f₁ : M → M'} {s t : Set M} {x : M} {m n : ℕ∞} + +section Module + +theorem DifferentiableWithinAt.comp_mdifferentiableWithinAt + {g : F → F'} {f : M → F} {s : Set M} {t : Set F} + {x : M} (hg : DifferentiableWithinAt 𝕜 g t (f x)) (hf : MDifferentiableWithinAt I 𝓘(𝕜, F) f s x) + (h : MapsTo f s t) : MDifferentiableWithinAt I 𝓘(𝕜, F') (g ∘ f) s x := + hg.mdifferentiableWithinAt.comp x hf h + +theorem DifferentiableAt.comp_mdifferentiableWithinAt {g : F → F'} {f : M → F} {s : Set M} + {x : M} (hg : DifferentiableAt 𝕜 g (f x)) (hf : MDifferentiableWithinAt I 𝓘(𝕜, F) f s x) : + MDifferentiableWithinAt I 𝓘(𝕜, F') (g ∘ f) s x := + hg.mdifferentiableAt.comp_mdifferentiableWithinAt x hf + +theorem DifferentiableAt.comp_mdifferentiableAt + {g : F → F'} {f : M → F} {x : M} (hg : DifferentiableAt 𝕜 g (f x)) + (hf : MDifferentiableAt I 𝓘(𝕜, F) f x) : MDifferentiableAt I 𝓘(𝕜, F') (g ∘ f) x := + hg.comp_mdifferentiableWithinAt hf + +theorem Differentiable.comp_mdifferentiableWithinAt {g : F → F'} {f : M → F} {s : Set M} {x : M} + (hg : Differentiable 𝕜 g) (hf : MDifferentiableWithinAt I 𝓘(𝕜, F) f s x) : + MDifferentiableWithinAt I 𝓘(𝕜, F') (g ∘ f) s x := + hg.differentiableAt.comp_mdifferentiableWithinAt hf + +theorem Differentiable.comp_mdifferentiableAt + {g : F → F'} {f : M → F} {x : M} (hg : Differentiable 𝕜 g) + (hf : MDifferentiableAt I 𝓘(𝕜, F) f x) : MDifferentiableAt I 𝓘(𝕜, F') (g ∘ f) x := + hg.comp_mdifferentiableWithinAt hf + +theorem Differentiable.comp_mdifferentiable {g : F → F'} {f : M → F} (hg : Differentiable 𝕜 g) + (hf : MDifferentiable I 𝓘(𝕜, F) f) : MDifferentiable I 𝓘(𝕜, F') (g ∘ f) := fun x => + hg.differentiableAt.comp_mdifferentiableAt (hf x) + +end Module + +/-! ### Linear maps between normed spaces are smooth -/ + +theorem MDifferentiableWithinAt.clm_precomp {f : M → F₁ →L[𝕜] F₂} {s : Set M} {x : M} + (hf : MDifferentiableWithinAt I 𝓘(𝕜, F₁ →L[𝕜] F₂) f s x) : + MDifferentiableWithinAt I 𝓘(𝕜, (F₂ →L[𝕜] F₃) →L[𝕜] (F₁ →L[𝕜] F₃)) + (fun y ↦ (f y).precomp F₃ : M → (F₂ →L[𝕜] F₃) →L[𝕜] (F₁ →L[𝕜] F₃)) s x := + Differentiable.comp_mdifferentiableWithinAt (g := (ContinuousLinearMap.compL 𝕜 F₁ F₂ F₃).flip) + (ContinuousLinearMap.compL 𝕜 F₁ F₂ F₃).flip.differentiable hf + +nonrec theorem MDifferentiableAt.clm_precomp {f : M → F₁ →L[𝕜] F₂} {x : M} + (hf : MDifferentiableAt I 𝓘(𝕜, F₁ →L[𝕜] F₂) f x) : + MDifferentiableAt I 𝓘(𝕜, (F₂ →L[𝕜] F₃) →L[𝕜] (F₁ →L[𝕜] F₃)) + (fun y ↦ (f y).precomp F₃ : M → (F₂ →L[𝕜] F₃) →L[𝕜] (F₁ →L[𝕜] F₃)) x := + Differentiable.comp_mdifferentiableAt (g := (ContinuousLinearMap.compL 𝕜 F₁ F₂ F₃).flip) + (ContinuousLinearMap.compL 𝕜 F₁ F₂ F₃).flip.differentiable hf + +theorem MDifferentiableOn.clm_precomp {f : M → F₁ →L[𝕜] F₂} {s : Set M} + (hf : MDifferentiableOn I 𝓘(𝕜, F₁ →L[𝕜] F₂) f s) : + MDifferentiableOn I 𝓘(𝕜, (F₂ →L[𝕜] F₃) →L[𝕜] (F₁ →L[𝕜] F₃)) + (fun y ↦ (f y).precomp F₃ : M → (F₂ →L[𝕜] F₃) →L[𝕜] (F₁ →L[𝕜] F₃)) s := fun x hx ↦ + (hf x hx).clm_precomp + +theorem MDifferentiable.clm_precomp + {f : M → F₁ →L[𝕜] F₂} (hf : MDifferentiable I 𝓘(𝕜, F₁ →L[𝕜] F₂) f) : + MDifferentiable I 𝓘(𝕜, (F₂ →L[𝕜] F₃) →L[𝕜] (F₁ →L[𝕜] F₃)) + (fun y ↦ (f y).precomp F₃ : M → (F₂ →L[𝕜] F₃) →L[𝕜] (F₁ →L[𝕜] F₃)) := fun x ↦ + (hf x).clm_precomp + +theorem MDifferentiableWithinAt.clm_postcomp {f : M → F₂ →L[𝕜] F₃} {s : Set M} {x : M} + (hf : MDifferentiableWithinAt I 𝓘(𝕜, F₂ →L[𝕜] F₃) f s x) : + MDifferentiableWithinAt I 𝓘(𝕜, (F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₃)) + (fun y ↦ (f y).postcomp F₁ : M → (F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₃)) s x := + Differentiable.comp_mdifferentiableWithinAt (F' := (F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₃)) + (g := ContinuousLinearMap.compL 𝕜 F₁ F₂ F₃) + (ContinuousLinearMap.compL 𝕜 F₁ F₂ F₃).differentiable hf + +theorem MDifferentiableAt.clm_postcomp {f : M → F₂ →L[𝕜] F₃} {x : M} + (hf : MDifferentiableAt I 𝓘(𝕜, F₂ →L[𝕜] F₃) f x) : + MDifferentiableAt I 𝓘(𝕜, (F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₃)) + (fun y ↦ (f y).postcomp F₁ : M → (F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₃)) x := + Differentiable.comp_mdifferentiableAt (F' := (F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₃)) + (g := ContinuousLinearMap.compL 𝕜 F₁ F₂ F₃) + (ContinuousLinearMap.compL 𝕜 F₁ F₂ F₃).differentiable hf + +nonrec theorem MDifferentiableOn.clm_postcomp {f : M → F₂ →L[𝕜] F₃} {s : Set M} + (hf : MDifferentiableOn I 𝓘(𝕜, F₂ →L[𝕜] F₃) f s) : + MDifferentiableOn I 𝓘(𝕜, (F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₃)) + (fun y ↦ (f y).postcomp F₁ : M → (F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₃)) s := fun x hx ↦ + (hf x hx).clm_postcomp + +theorem MDifferentiable.clm_postcomp + {f : M → F₂ →L[𝕜] F₃} (hf : MDifferentiable I 𝓘(𝕜, F₂ →L[𝕜] F₃) f) : + MDifferentiable I 𝓘(𝕜, (F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₃)) + (fun y ↦ (f y).postcomp F₁ : M → (F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₃)) := fun x ↦ + (hf x).clm_postcomp + +theorem MDifferentiableWithinAt.clm_comp + {g : M → F₁ →L[𝕜] F₃} {f : M → F₂ →L[𝕜] F₁} {s : Set M} {x : M} + (hg : MDifferentiableWithinAt I 𝓘(𝕜, F₁ →L[𝕜] F₃) g s x) + (hf : MDifferentiableWithinAt I 𝓘(𝕜, F₂ →L[𝕜] F₁) f s x) : + MDifferentiableWithinAt I 𝓘(𝕜, F₂ →L[𝕜] F₃) (fun x => (g x).comp (f x)) s x := + Differentiable.comp_mdifferentiableWithinAt + (g := fun x : (F₁ →L[𝕜] F₃) × (F₂ →L[𝕜] F₁) => x.1.comp x.2) + (f := fun x => (g x, f x)) (differentiable_fst.clm_comp differentiable_snd) + (hg.prod_mk_space hf) + +theorem MDifferentiableAt.clm_comp {g : M → F₁ →L[𝕜] F₃} {f : M → F₂ →L[𝕜] F₁} {x : M} + (hg : MDifferentiableAt I 𝓘(𝕜, F₁ →L[𝕜] F₃) g x) + (hf : MDifferentiableAt I 𝓘(𝕜, F₂ →L[𝕜] F₁) f x) : + MDifferentiableAt I 𝓘(𝕜, F₂ →L[𝕜] F₃) (fun x => (g x).comp (f x)) x := + (hg.mdifferentiableWithinAt.clm_comp hf.mdifferentiableWithinAt).mdifferentiableAt Filter.univ_mem + +theorem MDifferentiableOn.clm_comp {g : M → F₁ →L[𝕜] F₃} {f : M → F₂ →L[𝕜] F₁} {s : Set M} + (hg : MDifferentiableOn I 𝓘(𝕜, F₁ →L[𝕜] F₃) g s) + (hf : MDifferentiableOn I 𝓘(𝕜, F₂ →L[𝕜] F₁) f s) : + MDifferentiableOn I 𝓘(𝕜, F₂ →L[𝕜] F₃) (fun x => (g x).comp (f x)) s := fun x hx => + (hg x hx).clm_comp (hf x hx) + +theorem MDifferentiable.clm_comp {g : M → F₁ →L[𝕜] F₃} {f : M → F₂ →L[𝕜] F₁} + (hg : MDifferentiable I 𝓘(𝕜, F₁ →L[𝕜] F₃) g) (hf : MDifferentiable I 𝓘(𝕜, F₂ →L[𝕜] F₁) f) : + MDifferentiable I 𝓘(𝕜, F₂ →L[𝕜] F₃) fun x => (g x).comp (f x) := fun x => (hg x).clm_comp (hf x) + +/-- Applying a linear map to a vector is smooth within a set. Version in vector spaces. For a +version in nontrivial vector bundles, see `MDifferentiableWithinAt.clm_apply_of_inCoordinates`. -/ +theorem MDifferentiableWithinAt.clm_apply {g : M → F₁ →L[𝕜] F₂} {f : M → F₁} {s : Set M} {x : M} + (hg : MDifferentiableWithinAt I 𝓘(𝕜, F₁ →L[𝕜] F₂) g s x) + (hf : MDifferentiableWithinAt I 𝓘(𝕜, F₁) f s x) : + MDifferentiableWithinAt I 𝓘(𝕜, F₂) (fun x => g x (f x)) s x := + DifferentiableWithinAt.comp_mdifferentiableWithinAt (t := univ) + (g := fun x : (F₁ →L[𝕜] F₂) × F₁ => x.1 x.2) + (by apply (Differentiable.differentiableAt _).differentiableWithinAt + exact differentiable_fst.clm_apply differentiable_snd) (hg.prod_mk_space hf) + (by simp_rw [mapsTo_univ]) + +/-- Applying a linear map to a vector is smooth. Version in vector spaces. For a +version in nontrivial vector bundles, see `MDifferentiableAt.clm_apply_of_inCoordinates`. -/ +theorem MDifferentiableAt.clm_apply {g : M → F₁ →L[𝕜] F₂} {f : M → F₁} {x : M} + (hg : MDifferentiableAt I 𝓘(𝕜, F₁ →L[𝕜] F₂) g x) (hf : MDifferentiableAt I 𝓘(𝕜, F₁) f x) : + MDifferentiableAt I 𝓘(𝕜, F₂) (fun x => g x (f x)) x := + DifferentiableWithinAt.comp_mdifferentiableWithinAt (t := univ) + (g := fun x : (F₁ →L[𝕜] F₂) × F₁ => x.1 x.2) + (by apply (Differentiable.differentiableAt _).differentiableWithinAt + exact differentiable_fst.clm_apply differentiable_snd) (hg.prod_mk_space hf) + (by simp_rw [mapsTo_univ]) + +theorem MDifferentiableOn.clm_apply {g : M → F₁ →L[𝕜] F₂} {f : M → F₁} {s : Set M} + (hg : MDifferentiableOn I 𝓘(𝕜, F₁ →L[𝕜] F₂) g s) (hf : MDifferentiableOn I 𝓘(𝕜, F₁) f s) : + MDifferentiableOn I 𝓘(𝕜, F₂) (fun x => g x (f x)) s := fun x hx => (hg x hx).clm_apply (hf x hx) + +theorem MDifferentiable.clm_apply {g : M → F₁ →L[𝕜] F₂} {f : M → F₁} + (hg : MDifferentiable I 𝓘(𝕜, F₁ →L[𝕜] F₂) g) (hf : MDifferentiable I 𝓘(𝕜, F₁) f) : + MDifferentiable I 𝓘(𝕜, F₂) fun x => g x (f x) := fun x => (hg x).clm_apply (hf x) + +theorem MDifferentiableWithinAt.cle_arrowCongr {f : M → F₁ ≃L[𝕜] F₂} {g : M → F₃ ≃L[𝕜] F₄} + {s : Set M} {x : M} + (hf : MDifferentiableWithinAt I 𝓘(𝕜, F₂ →L[𝕜] F₁) (fun x ↦ ((f x).symm : F₂ →L[𝕜] F₁)) s x) + (hg : MDifferentiableWithinAt I 𝓘(𝕜, F₃ →L[𝕜] F₄) (fun x ↦ (g x : F₃ →L[𝕜] F₄)) s x) : + MDifferentiableWithinAt I 𝓘(𝕜, (F₁ →L[𝕜] F₃) →L[𝕜] (F₂ →L[𝕜] F₄)) + (fun y ↦ (f y).arrowCongr (g y) : M → (F₁ →L[𝕜] F₃) →L[𝕜] (F₂ →L[𝕜] F₄)) s x := + show MDifferentiableWithinAt I 𝓘(𝕜, (F₁ →L[𝕜] F₃) →L[𝕜] (F₂ →L[𝕜] F₄)) + (fun y ↦ (((f y).symm : F₂ →L[𝕜] F₁).precomp F₄).comp ((g y : F₃ →L[𝕜] F₄).postcomp F₁)) s x + from hf.clm_precomp (F₃ := F₄) |>.clm_comp <| hg.clm_postcomp (F₁ := F₁) + +theorem MDifferentiableAt.cle_arrowCongr {f : M → F₁ ≃L[𝕜] F₂} {g : M → F₃ ≃L[𝕜] F₄} {x : M} + (hf : MDifferentiableAt I 𝓘(𝕜, F₂ →L[𝕜] F₁) (fun x ↦ ((f x).symm : F₂ →L[𝕜] F₁)) x) + (hg : MDifferentiableAt I 𝓘(𝕜, F₃ →L[𝕜] F₄) (fun x ↦ (g x : F₃ →L[𝕜] F₄)) x) : + MDifferentiableAt I 𝓘(𝕜, (F₁ →L[𝕜] F₃) →L[𝕜] (F₂ →L[𝕜] F₄)) + (fun y ↦ (f y).arrowCongr (g y) : M → (F₁ →L[𝕜] F₃) →L[𝕜] (F₂ →L[𝕜] F₄)) x := + show MDifferentiableAt I 𝓘(𝕜, (F₁ →L[𝕜] F₃) →L[𝕜] (F₂ →L[𝕜] F₄)) + (fun y ↦ (((f y).symm : F₂ →L[𝕜] F₁).precomp F₄).comp ((g y : F₃ →L[𝕜] F₄).postcomp F₁)) x + from hf.clm_precomp (F₃ := F₄) |>.clm_comp <| hg.clm_postcomp (F₁ := F₁) + +theorem MDifferentiableOn.cle_arrowCongr {f : M → F₁ ≃L[𝕜] F₂} {g : M → F₃ ≃L[𝕜] F₄} {s : Set M} + (hf : MDifferentiableOn I 𝓘(𝕜, F₂ →L[𝕜] F₁) (fun x ↦ ((f x).symm : F₂ →L[𝕜] F₁)) s) + (hg : MDifferentiableOn I 𝓘(𝕜, F₃ →L[𝕜] F₄) (fun x ↦ (g x : F₃ →L[𝕜] F₄)) s) : + MDifferentiableOn I 𝓘(𝕜, (F₁ →L[𝕜] F₃) →L[𝕜] (F₂ →L[𝕜] F₄)) + (fun y ↦ (f y).arrowCongr (g y) : M → (F₁ →L[𝕜] F₃) →L[𝕜] (F₂ →L[𝕜] F₄)) s := fun x hx ↦ + (hf x hx).cle_arrowCongr (hg x hx) + +theorem MDifferentiable.cle_arrowCongr {f : M → F₁ ≃L[𝕜] F₂} {g : M → F₃ ≃L[𝕜] F₄} + (hf : MDifferentiable I 𝓘(𝕜, F₂ →L[𝕜] F₁) (fun x ↦ ((f x).symm : F₂ →L[𝕜] F₁))) + (hg : MDifferentiable I 𝓘(𝕜, F₃ →L[𝕜] F₄) (fun x ↦ (g x : F₃ →L[𝕜] F₄))) : + MDifferentiable I 𝓘(𝕜, (F₁ →L[𝕜] F₃) →L[𝕜] (F₂ →L[𝕜] F₄)) + (fun y ↦ (f y).arrowCongr (g y) : M → (F₁ →L[𝕜] F₃) →L[𝕜] (F₂ →L[𝕜] F₄)) := fun x ↦ + (hf x).cle_arrowCongr (hg x) + +theorem MDifferentiableWithinAt.clm_prodMap {g : M → F₁ →L[𝕜] F₃} {f : M → F₂ →L[𝕜] F₄} {s : Set M} + {x : M} (hg : MDifferentiableWithinAt I 𝓘(𝕜, F₁ →L[𝕜] F₃) g s x) + (hf : MDifferentiableWithinAt I 𝓘(𝕜, F₂ →L[𝕜] F₄) f s x) : + MDifferentiableWithinAt I 𝓘(𝕜, F₁ × F₂ →L[𝕜] F₃ × F₄) (fun x => (g x).prodMap (f x)) s x := + Differentiable.comp_mdifferentiableWithinAt + (g := fun x : (F₁ →L[𝕜] F₃) × (F₂ →L[𝕜] F₄) => x.1.prodMap x.2) + (f := fun x => (g x, f x)) (ContinuousLinearMap.prodMapL 𝕜 F₁ F₃ F₂ F₄).differentiable + (hg.prod_mk_space hf) + +nonrec theorem MDifferentiableAt.clm_prodMap {g : M → F₁ →L[𝕜] F₃} {f : M → F₂ →L[𝕜] F₄} {x : M} + (hg : MDifferentiableAt I 𝓘(𝕜, F₁ →L[𝕜] F₃) g x) + (hf : MDifferentiableAt I 𝓘(𝕜, F₂ →L[𝕜] F₄) f x) : + MDifferentiableAt I 𝓘(𝕜, F₁ × F₂ →L[𝕜] F₃ × F₄) (fun x => (g x).prodMap (f x)) x := + Differentiable.comp_mdifferentiableWithinAt + (g := fun x : (F₁ →L[𝕜] F₃) × (F₂ →L[𝕜] F₄) => x.1.prodMap x.2) + (f := fun x => (g x, f x)) (ContinuousLinearMap.prodMapL 𝕜 F₁ F₃ F₂ F₄).differentiable + (hg.prod_mk_space hf) + +theorem MDifferentiableOn.clm_prodMap {g : M → F₁ →L[𝕜] F₃} {f : M → F₂ →L[𝕜] F₄} {s : Set M} + (hg : MDifferentiableOn I 𝓘(𝕜, F₁ →L[𝕜] F₃) g s) + (hf : MDifferentiableOn I 𝓘(𝕜, F₂ →L[𝕜] F₄) f s) : + MDifferentiableOn I 𝓘(𝕜, F₁ × F₂ →L[𝕜] F₃ × F₄) (fun x => (g x).prodMap (f x)) s := fun x hx => + (hg x hx).clm_prodMap (hf x hx) + +theorem MDifferentiable.clm_prodMap {g : M → F₁ →L[𝕜] F₃} {f : M → F₂ →L[𝕜] F₄} + (hg : MDifferentiable I 𝓘(𝕜, F₁ →L[𝕜] F₃) g) (hf : MDifferentiable I 𝓘(𝕜, F₂ →L[𝕜] F₄) f) : + MDifferentiable I 𝓘(𝕜, F₁ × F₂ →L[𝕜] F₃ × F₄) fun x => (g x).prodMap (f x) := fun x => + (hg x).clm_prodMap (hf x) + +/-! ### Smoothness of scalar multiplication -/ + +variable {V : Type*} [NormedAddCommGroup V] [NormedSpace 𝕜 V] + +theorem MDifferentiableWithinAt.smul {f : M → 𝕜} {g : M → V} + (hf : MDifferentiableWithinAt I 𝓘(𝕜) f s x) (hg : MDifferentiableWithinAt I 𝓘(𝕜, V) g s x) : + MDifferentiableWithinAt I 𝓘(𝕜, V) (fun p => f p • g p) s x := + ((contMDiff_smul.of_le le_top).mdifferentiable le_rfl _).comp_mdifferentiableWithinAt x + (hf.prod_mk hg) + +theorem MDifferentiableAt.smul {f : M → 𝕜} {g : M → V} (hf : MDifferentiableAt I 𝓘(𝕜) f x) + (hg : MDifferentiableAt I 𝓘(𝕜, V) g x) : MDifferentiableAt I 𝓘(𝕜, V) (fun p => f p • g p) x := + ((contMDiff_smul.of_le le_top).mdifferentiable le_rfl _).comp x (hf.prod_mk hg) + +theorem MDifferentiableOn.smul {f : M → 𝕜} {g : M → V} (hf : MDifferentiableOn I 𝓘(𝕜) f s) + (hg : MDifferentiableOn I 𝓘(𝕜, V) g s) : MDifferentiableOn I 𝓘(𝕜, V) (fun p => f p • g p) s := + fun x hx => (hf x hx).smul (hg x hx) + +theorem MDifferentiable.smul {f : M → 𝕜} {g : M → V} (hf : MDifferentiable I 𝓘(𝕜) f) + (hg : MDifferentiable I 𝓘(𝕜, V) g) : MDifferentiable I 𝓘(𝕜, V) fun p => f p • g p := fun x => + (hf x).smul (hg x) From 0b0b3f79dcc8aa88fd51f34904d6b56af1de1e45 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Thu, 28 Nov 2024 15:29:13 +0000 Subject: [PATCH 387/829] feat: the pullback of a vector field in a vector space (#19357) We prove notably that the pullback commutes with the Lie bracket. --- Mathlib/Analysis/Calculus/VectorField.lean | 230 ++++++++++++++++++++- 1 file changed, 229 insertions(+), 1 deletion(-) diff --git a/Mathlib/Analysis/Calculus/VectorField.lean b/Mathlib/Analysis/Calculus/VectorField.lean index ac4eb32fdc707..0581fb044d1d1 100644 --- a/Mathlib/Analysis/Calculus/VectorField.lean +++ b/Mathlib/Analysis/Calculus/VectorField.lean @@ -16,7 +16,13 @@ Notably, we define the pullback of a vector field under a map, as `VectorField.pullback 𝕜 f V x := (fderiv 𝕜 f x).inverse (V (f x))` (together with the same notion within a set). -In addition to comprehensive API on this notion, the main result is the following: +We also define the Lie bracket of two vector fields as +`VectorField.lieBracket 𝕜 V W x := fderiv 𝕜 W x (V x) - fderiv 𝕜 V x (W x)` +(together with the same notion within a set). + +In addition to comprehensive API on these two notions, the main results are the following: +* `VectorField.pullback_lieBracket` states that the pullback of the Lie bracket + is the Lie bracket of the pullbacks, when the second derivative is symmetric. * `VectorField.leibniz_identity_lieBracket` is the Leibniz identity `[U, [V, W]] = [[U, V], W] + [V, [U, W]]`. @@ -340,4 +346,226 @@ lemma leibniz_identity_lieBracket [IsRCLikeNormedField 𝕜] {U V W : E → E} { simp only [← lieBracketWithin_univ, ← contDiffWithinAt_univ] at hU hV hW ⊢ exact leibniz_identity_lieBracketWithin uniqueDiffOn_univ (by simp) (mem_univ _) hU hV hW + +/-! +### The pullback of vector fields in a vector space +-/ + +variable (𝕜) in +/-- The pullback of a vector field under a function, defined +as `(f^* V) (x) = Df(x)^{-1} (V (f x))`. If `Df(x)` is not invertible, we use the junk value `0`. +-/ +def pullback (f : E → F) (V : F → F) (x : E) : E := (fderiv 𝕜 f x).inverse (V (f x)) + +variable (𝕜) in +/-- The pullback within a set of a vector field under a function, defined +as `(f^* V) (x) = Df(x)^{-1} (V (f x))` where `Df(x)` is the derivative of `f` within `s`. +If `Df(x)` is not invertible, we use the junk value `0`. +-/ +def pullbackWithin (f : E → F) (V : F → F) (s : Set E) (x : E) : E := + (fderivWithin 𝕜 f s x).inverse (V (f x)) + +lemma pullbackWithin_eq {f : E → F} {V : F → F} {s : Set E} : + pullbackWithin 𝕜 f V s = fun x ↦ (fderivWithin 𝕜 f s x).inverse (V (f x)) := rfl + +lemma pullback_eq_of_fderiv_eq + {f : E → F} {M : E ≃L[𝕜] F} {x : E} (hf : M = fderiv 𝕜 f x) (V : F → F) : + pullback 𝕜 f V x = M.symm (V (f x)) := by + simp [pullback, ← hf] + +lemma pullback_eq_of_not_isInvertible {f : E → F} {x : E} + (h : ¬(fderiv 𝕜 f x).IsInvertible) (V : F → F) : + pullback 𝕜 f V x = 0 := by + simp [pullback, h] + +lemma pullbackWithin_eq_of_not_isInvertible {f : E → F} {x : E} + (h : ¬(fderivWithin 𝕜 f s x).IsInvertible) (V : F → F) : + pullbackWithin 𝕜 f V s x = 0 := by + simp [pullbackWithin, h] + +lemma pullbackWithin_eq_of_fderivWithin_eq + {f : E → F} {M : E ≃L[𝕜] F} {x : E} (hf : M = fderivWithin 𝕜 f s x) (V : F → F) : + pullbackWithin 𝕜 f V s x = M.symm (V (f x)) := by + simp [pullbackWithin, ← hf] + +@[simp] lemma pullbackWithin_univ {f : E → F} {V : F → F} : + pullbackWithin 𝕜 f V univ = pullback 𝕜 f V := by + ext x + simp [pullbackWithin, pullback] + +open scoped Topology Filter + +lemma fderiv_pullback (f : E → F) (V : F → F) (x : E) (h'f : (fderiv 𝕜 f x).IsInvertible) : + fderiv 𝕜 f x (pullback 𝕜 f V x) = V (f x) := by + rcases h'f with ⟨M, hM⟩ + simp [pullback_eq_of_fderiv_eq hM, ← hM] + +lemma fderivWithin_pullbackWithin {f : E → F} {V : F → F} {x : E} + (h'f : (fderivWithin 𝕜 f s x).IsInvertible) : + fderivWithin 𝕜 f s x (pullbackWithin 𝕜 f V s x) = V (f x) := by + rcases h'f with ⟨M, hM⟩ + simp [pullbackWithin_eq_of_fderivWithin_eq hM, ← hM] + +open Set + +variable [CompleteSpace E] + +/-- If a `C^2` map has an invertible derivative within a set at a point, then nearby derivatives +can be written as continuous linear equivs, which depend in a `C^1` way on the point, as well as +their inverse, and moreover one can compute the derivative of the inverse. -/ +lemma _root_.exists_continuousLinearEquiv_fderivWithin_symm_eq + {f : E → F} {s : Set E} {x : E} (h'f : ContDiffWithinAt 𝕜 2 f s x) + (hf : (fderivWithin 𝕜 f s x).IsInvertible) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) : + ∃ N : E → (E ≃L[𝕜] F), ContDiffWithinAt 𝕜 1 (fun y ↦ (N y : E →L[𝕜] F)) s x + ∧ ContDiffWithinAt 𝕜 1 (fun y ↦ ((N y).symm : F →L[𝕜] E)) s x + ∧ (∀ᶠ y in 𝓝[s] x, N y = fderivWithin 𝕜 f s y) + ∧ ∀ v, fderivWithin 𝕜 (fun y ↦ ((N y).symm : F →L[𝕜] E)) s x v + = - (N x).symm ∘L ((fderivWithin 𝕜 (fderivWithin 𝕜 f s) s x v)) ∘L (N x).symm := by + classical + rcases hf with ⟨M, hM⟩ + let U := {y | ∃ (N : E ≃L[𝕜] F), N = fderivWithin 𝕜 f s y} + have hU : U ∈ 𝓝[s] x := by + have I : range ((↑) : (E ≃L[𝕜] F) → E →L[𝕜] F) ∈ 𝓝 (fderivWithin 𝕜 f s x) := by + rw [← hM] + exact M.nhds + have : ContinuousWithinAt (fderivWithin 𝕜 f s) s x := + (h'f.fderivWithin_right (m := 1) hs le_rfl hx).continuousWithinAt + exact this I + let N : E → (E ≃L[𝕜] F) := fun x ↦ if h : x ∈ U then h.choose else M + have eN : (fun y ↦ (N y : E →L[𝕜] F)) =ᶠ[𝓝[s] x] fun y ↦ fderivWithin 𝕜 f s y := by + filter_upwards [hU] with y hy + simpa only [hy, ↓reduceDIte, N] using Exists.choose_spec hy + have e'N : N x = fderivWithin 𝕜 f s x := by apply mem_of_mem_nhdsWithin hx eN + have hN : ContDiffWithinAt 𝕜 1 (fun y ↦ (N y : E →L[𝕜] F)) s x := by + have : ContDiffWithinAt 𝕜 1 (fun y ↦ fderivWithin 𝕜 f s y) s x := + h'f.fderivWithin_right (m := 1) hs le_rfl hx + apply this.congr_of_eventuallyEq eN e'N + have hN' : ContDiffWithinAt 𝕜 1 (fun y ↦ ((N y).symm : F →L[𝕜] E)) s x := by + have : ContDiffWithinAt 𝕜 1 (ContinuousLinearMap.inverse ∘ (fun y ↦ (N y : E →L[𝕜] F))) s x := + (contDiffAt_map_inverse (N x)).comp_contDiffWithinAt x hN + convert this with y + simp only [Function.comp_apply, ContinuousLinearMap.inverse_equiv] + refine ⟨N, hN, hN', eN, fun v ↦ ?_⟩ + have A' y : ContinuousLinearMap.compL 𝕜 F E F (N y : E →L[𝕜] F) ((N y).symm : F →L[𝕜] E) + = ContinuousLinearMap.id 𝕜 F := by ext; simp + have : fderivWithin 𝕜 (fun y ↦ ContinuousLinearMap.compL 𝕜 F E F (N y : E →L[𝕜] F) + ((N y).symm : F →L[𝕜] E)) s x v = 0 := by + simp [A', fderivWithin_const_apply, hs x hx] + have I : (N x : E →L[𝕜] F) ∘L (fderivWithin 𝕜 (fun y ↦ ((N y).symm : F →L[𝕜] E)) s x v) = + - (fderivWithin 𝕜 (fun y ↦ (N y : E →L[𝕜] F)) s x v) ∘L ((N x).symm : F →L[𝕜] E) := by + rw [ContinuousLinearMap.fderivWithin_of_bilinear _ (hN.differentiableWithinAt le_rfl) + (hN'.differentiableWithinAt le_rfl) (hs x hx)] at this + simpa [eq_neg_iff_add_eq_zero] using this + have B (M : F →L[𝕜] E) : M = ((N x).symm : F →L[𝕜] E) ∘L ((N x) ∘L M) := by + ext; simp + rw [B (fderivWithin 𝕜 (fun y ↦ ((N y).symm : F →L[𝕜] E)) s x v), I] + simp only [ContinuousLinearMap.comp_neg, neg_inj, eN.fderivWithin_eq e'N] + +lemma DifferentiableWithinAt.pullbackWithin {f : E → F} {V : F → F} {s : Set E} {t : Set F} {x : E} + (hV : DifferentiableWithinAt 𝕜 V t (f x)) + (hf : ContDiffWithinAt 𝕜 2 f s x) (hf' : (fderivWithin 𝕜 f s x).IsInvertible) + (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) (hst : MapsTo f s t) : + DifferentiableWithinAt 𝕜 (pullbackWithin 𝕜 f V s) s x := by + rcases exists_continuousLinearEquiv_fderivWithin_symm_eq hf hf' hs hx + with ⟨M, -, M_symm_smooth, hM, -⟩ + simp only [pullbackWithin_eq] + have : DifferentiableWithinAt 𝕜 (fun y ↦ ((M y).symm : F →L[𝕜] E) (V (f y))) s x := by + apply DifferentiableWithinAt.clm_apply + · exact M_symm_smooth.differentiableWithinAt le_rfl + · exact hV.comp _ (hf.differentiableWithinAt one_le_two) hst + apply this.congr_of_eventuallyEq + · filter_upwards [hM] with y hy using by simp [← hy] + · have hMx : M x = fderivWithin 𝕜 f s x := by apply mem_of_mem_nhdsWithin hx hM + simp [← hMx] + +/-- If a `C^2` map has an invertible derivative at a point, then nearby derivatives can be written +as continuous linear equivs, which depend in a `C^1` way on the point, as well as their inverse, and +moreover one can compute the derivative of the inverse. -/ +lemma _root_.exists_continuousLinearEquiv_fderiv_symm_eq + {f : E → F} {x : E} (h'f : ContDiffAt 𝕜 2 f x) (hf : (fderiv 𝕜 f x).IsInvertible) : + ∃ N : E → (E ≃L[𝕜] F), ContDiffAt 𝕜 1 (fun y ↦ (N y : E →L[𝕜] F)) x + ∧ ContDiffAt 𝕜 1 (fun y ↦ ((N y).symm : F →L[𝕜] E)) x + ∧ (∀ᶠ y in 𝓝 x, N y = fderiv 𝕜 f y) + ∧ ∀ v, fderiv 𝕜 (fun y ↦ ((N y).symm : F →L[𝕜] E)) x v + = - (N x).symm ∘L ((fderiv 𝕜 (fderiv 𝕜 f) x v)) ∘L (N x).symm := by + simp only [← fderivWithin_univ, ← contDiffWithinAt_univ, ← nhdsWithin_univ] at hf h'f ⊢ + exact exists_continuousLinearEquiv_fderivWithin_symm_eq h'f hf uniqueDiffOn_univ (mem_univ _) + +/-- The Lie bracket commutes with taking pullbacks. This requires the function to have symmetric +second derivative. Version in a complete space. One could also give a version avoiding +completeness but requiring that `f` is a local diffeo. -/ +lemma pullbackWithin_lieBracketWithin_of_isSymmSndFDerivWithinAt + {f : E → F} {V W : F → F} {x : E} {t : Set F} + (hf : IsSymmSndFDerivWithinAt 𝕜 f s x) (h'f : ContDiffWithinAt 𝕜 2 f s x) + (hV : DifferentiableWithinAt 𝕜 V t (f x)) (hW : DifferentiableWithinAt 𝕜 W t (f x)) + (hu : UniqueDiffOn 𝕜 s) (hx : x ∈ s) (hst : MapsTo f s t) : + pullbackWithin 𝕜 f (lieBracketWithin 𝕜 V W t) s x + = lieBracketWithin 𝕜 (pullbackWithin 𝕜 f V s) (pullbackWithin 𝕜 f W s) s x := by + by_cases h : (fderivWithin 𝕜 f s x).IsInvertible; swap + · simp [pullbackWithin_eq_of_not_isInvertible h, lieBracketWithin_eq] + rcases exists_continuousLinearEquiv_fderivWithin_symm_eq h'f h hu hx + with ⟨M, -, M_symm_smooth, hM, M_diff⟩ + have hMx : M x = fderivWithin 𝕜 f s x := (mem_of_mem_nhdsWithin hx hM :) + have AV : fderivWithin 𝕜 (pullbackWithin 𝕜 f V s) s x = + fderivWithin 𝕜 (fun y ↦ ((M y).symm : F →L[𝕜] E) (V (f y))) s x := by + apply Filter.EventuallyEq.fderivWithin_eq_of_mem _ hx + filter_upwards [hM] with y hy using pullbackWithin_eq_of_fderivWithin_eq hy _ + have AW : fderivWithin 𝕜 (pullbackWithin 𝕜 f W s) s x = + fderivWithin 𝕜 (fun y ↦ ((M y).symm : F →L[𝕜] E) (W (f y))) s x := by + apply Filter.EventuallyEq.fderivWithin_eq_of_mem _ hx + filter_upwards [hM] with y hy using pullbackWithin_eq_of_fderivWithin_eq hy _ + have Af : DifferentiableWithinAt 𝕜 f s x := h'f.differentiableWithinAt one_le_two + simp only [lieBracketWithin_eq, pullbackWithin_eq_of_fderivWithin_eq hMx, map_sub, AV, AW] + rw [fderivWithin_clm_apply, fderivWithin_clm_apply] + · simp [fderivWithin_comp' x hW Af hst (hu x hx), ← hMx, + fderivWithin_comp' x hV Af hst (hu x hx), M_diff, hf.eq] + · exact hu x hx + · exact M_symm_smooth.differentiableWithinAt le_rfl + · exact hV.comp x Af hst + · exact hu x hx + · exact M_symm_smooth.differentiableWithinAt le_rfl + · exact hW.comp x Af hst + +/-- The Lie bracket commutes with taking pullbacks. This requires the function to have symmetric +second derivative. Version in a complete space. One could also give a version avoiding +completeness but requiring that `f` is a local diffeo. Variant where unique differentiability and +the invariance property are only required in a smaller set `u`. -/ +lemma pullbackWithin_lieBracketWithin_of_isSymmSndFDerivWithinAt_of_eventuallyEq + {f : E → F} {V W : F → F} {x : E} {t : Set F} {u : Set E} + (hf : IsSymmSndFDerivWithinAt 𝕜 f s x) (h'f : ContDiffWithinAt 𝕜 2 f s x) + (hV : DifferentiableWithinAt 𝕜 V t (f x)) (hW : DifferentiableWithinAt 𝕜 W t (f x)) + (hu : UniqueDiffOn 𝕜 u) (hx : x ∈ u) (hst : MapsTo f u t) (hus : u =ᶠ[𝓝 x] s) : + pullbackWithin 𝕜 f (lieBracketWithin 𝕜 V W t) s x + = lieBracketWithin 𝕜 (pullbackWithin 𝕜 f V s) (pullbackWithin 𝕜 f W s) s x := calc + pullbackWithin 𝕜 f (lieBracketWithin 𝕜 V W t) s x + _ = pullbackWithin 𝕜 f (lieBracketWithin 𝕜 V W t) u x := by + simp only [pullbackWithin] + congr 2 + exact fderivWithin_congr_set hus.symm + _ = lieBracketWithin 𝕜 (pullbackWithin 𝕜 f V u) (pullbackWithin 𝕜 f W u) u x := + pullbackWithin_lieBracketWithin_of_isSymmSndFDerivWithinAt + (hf.congr_set hus.symm) (h'f.congr_set hus.symm) hV hW hu hx hst + _ = lieBracketWithin 𝕜 (pullbackWithin 𝕜 f V s) (pullbackWithin 𝕜 f W s) u x := by + apply Filter.EventuallyEq.lieBracketWithin_vectorField_eq_of_mem _ _ hx + · apply nhdsWithin_le_nhds + filter_upwards [fderivWithin_eventually_congr_set (𝕜 := 𝕜) (f := f) hus] with y hy + simp [pullbackWithin, hy] + · apply nhdsWithin_le_nhds + filter_upwards [fderivWithin_eventually_congr_set (𝕜 := 𝕜) (f := f) hus] with y hy + simp [pullbackWithin, hy] + _ = lieBracketWithin 𝕜 (pullbackWithin 𝕜 f V s) (pullbackWithin 𝕜 f W s) s x := + lieBracketWithin_congr_set hus + +/-- The Lie bracket commutes with taking pullbacks. This requires the function to have symmetric +second derivative. Version in a complete space. One could also give a version avoiding +completeness but requiring that `f` is a local diffeo. -/ +lemma pullback_lieBracket_of_isSymmSndFDerivAt {f : E → F} {V W : F → F} {x : E} + (hf : IsSymmSndFDerivAt 𝕜 f x) (h'f : ContDiffAt 𝕜 2 f x) + (hV : DifferentiableAt 𝕜 V (f x)) (hW : DifferentiableAt 𝕜 W (f x)) : + pullback 𝕜 f (lieBracket 𝕜 V W) x = lieBracket 𝕜 (pullback 𝕜 f V) (pullback 𝕜 f W) x := by + simp only [← lieBracketWithin_univ, ← pullbackWithin_univ, ← isSymmSndFDerivWithinAt_univ, + ← differentiableWithinAt_univ] at hf h'f hV hW ⊢ + exact pullbackWithin_lieBracketWithin_of_isSymmSndFDerivWithinAt hf h'f hV hW uniqueDiffOn_univ + (mem_univ _) (mapsTo_univ _ _) + end VectorField From 746714471ae87031bebd4bcdbc49f2aae01cae8f Mon Sep 17 00:00:00 2001 From: Justus Springer Date: Thu, 28 Nov 2024 15:29:14 +0000 Subject: [PATCH 388/829] refactor(Algebra/Polynomial/Laurent): simplify some proofs (#19534) This PR simplifies a few proofs about Laurent polynomials and adds the trivial theorem `Polynomial.toLaurent_eq_zero`. It also changes the type of `Lauren.isLocalization` to use `Submonoid.powers` instead of `Submonoid.closure`, which I think is the more natural thing to do here. Co-authored-by: Justus Springer <50165510+justus-springer@users.noreply.github.com> --- Mathlib/Algebra/Polynomial/Laurent.lean | 41 ++++++++++++------------- 1 file changed, 19 insertions(+), 22 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/Laurent.lean b/Mathlib/Algebra/Polynomial/Laurent.lean index 50464db951265..36ce911b4154b 100644 --- a/Mathlib/Algebra/Polynomial/Laurent.lean +++ b/Mathlib/Algebra/Polynomial/Laurent.lean @@ -332,8 +332,12 @@ theorem _root_.Polynomial.toLaurent_injective : theorem _root_.Polynomial.toLaurent_inj (f g : R[X]) : toLaurent f = toLaurent g ↔ f = g := ⟨fun h => Polynomial.toLaurent_injective h, congr_arg _⟩ -theorem _root_.Polynomial.toLaurent_ne_zero {f : R[X]} : f ≠ 0 ↔ toLaurent f ≠ 0 := - (map_ne_zero_iff _ Polynomial.toLaurent_injective).symm +theorem _root_.Polynomial.toLaurent_ne_zero {f : R[X]} : toLaurent f ≠ 0 ↔ f ≠ 0 := + map_ne_zero_iff _ Polynomial.toLaurent_injective + +@[simp] +theorem _root_.Polynomial.toLaurent_eq_zero {f : R[X]} : toLaurent f = 0 ↔ f = 0 := + map_eq_zero_iff _ Polynomial.toLaurent_injective theorem exists_T_pow (f : R[T;T⁻¹]) : ∃ (n : ℕ) (f' : R[X]), toLaurent f' = f * T n := by refine f.induction_on' ?_ fun n a => ?_ <;> clear f @@ -387,9 +391,9 @@ theorem toLaurent_support (f : R[X]) : f.toLaurent.support = f.support.map Nat.c generalize hd : f.support = s revert f refine Finset.induction_on s ?_ ?_ <;> clear s - · simp +contextual only [Polynomial.support_eq_empty, map_zero, - Finsupp.support_zero, eq_self_iff_true, imp_true_iff, Finset.map_empty, - Finsupp.support_eq_empty] + · intro f hf + rw [Finset.map_empty, Finsupp.support_eq_empty, toLaurent_eq_zero] + exact Polynomial.support_eq_empty.mp hf · intro a s as hf f fs have : (erase a f).toLaurent.support = s.map Nat.castEmbedding := by refine hf (f.erase a) ?_ @@ -421,21 +425,14 @@ theorem degree_eq_bot_iff {f : R[T;T⁻¹]} : f.degree = ⊥ ↔ f = 0 := by refine ⟨fun h => ?_, fun h => by rw [h, degree_zero]⟩ rw [degree, Finset.max_eq_sup_withBot] at h ext n - refine not_not.mp fun f0 => ?_ simp_rw [Finset.sup_eq_bot_iff, Finsupp.mem_support_iff, Ne, WithBot.coe_ne_bot] at h - exact h n f0 + exact not_not.mp (h n) section ExactDegrees @[simp] theorem degree_C_mul_T (n : ℤ) (a : R) (a0 : a ≠ 0) : degree (C a * T n) = n := by - rw [degree] - -- Porting note: was `convert Finset.max_singleton` - have : Finsupp.support (C a * T n) = {n} := by - refine support_eq_singleton.mpr ?_ - rw [← single_eq_C_mul_T] - simp only [single_eq_same, a0, Ne, not_false_iff, eq_self_iff_true, and_self_iff] - rw [this] + rw [degree, support_C_mul_T_of_ne_zero a0 n] exact Finset.max_singleton theorem degree_C_mul_T_ite [DecidableEq R] (n : ℤ) (a : R) : @@ -499,17 +496,17 @@ theorem algebraMap_X_pow (n : ℕ) : algebraMap R[X] R[T;T⁻¹] (X ^ n) = T n : theorem algebraMap_eq_toLaurent (f : R[X]) : algebraMap R[X] R[T;T⁻¹] f = toLaurent f := rfl -theorem isLocalization : IsLocalization (Submonoid.closure ({X} : Set R[X])) R[T;T⁻¹] := - { map_units' := fun t => by - cases' t with t ht - rcases Submonoid.mem_closure_singleton.mp ht with ⟨n, rfl⟩ - simp only [isUnit_T n, algebraMap_eq_toLaurent, Polynomial.toLaurent_X_pow] +theorem isLocalization : IsLocalization (Submonoid.powers (X : R[X])) R[T;T⁻¹] := + { map_units' := fun ⟨t, ht⟩ => by + obtain ⟨n, rfl⟩ := ht + rw [algebraMap_eq_toLaurent, toLaurent_X_pow] + exact isUnit_T ↑n surj' := fun f => by induction' f using LaurentPolynomial.induction_on_mul_T with f n - have := (Submonoid.closure ({X} : Set R[X])).pow_mem Submonoid.mem_closure_singleton_self n + have : X ^ n ∈ Submonoid.powers (X : R[X]) := ⟨n, rfl⟩ refine ⟨(f, ⟨_, this⟩), ?_⟩ - simp only [algebraMap_eq_toLaurent, Polynomial.toLaurent_X_pow, mul_T_assoc, - neg_add_cancel, T_zero, mul_one] + simp only [algebraMap_eq_toLaurent, toLaurent_X_pow, mul_T_assoc, neg_add_cancel, T_zero, + mul_one] exists_of_eq := fun {f g} => by rw [algebraMap_eq_toLaurent, algebraMap_eq_toLaurent, Polynomial.toLaurent_inj] rintro rfl From e71723be53d5b233ee010206c1ec451807f992ba Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 28 Nov 2024 16:12:25 +0000 Subject: [PATCH 389/829] feat: Goursat's lemma for subgroups (#18823) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Proves Goursat's lemma for subgroups. If `I` is a subgroup of `G × H` which projects fully on both factors, then there exist normal subgroups `G' ≤ G` and `H' ≤ H` such that `G' × H' ≤ I` and the image of `I` in `G ⧸ G' × H ⧸ H'` is the graph of an isomorphism `G ⧸ G' ≃ H ⧸ H'`. We construct `G'` and `H'` as explicit subgroups. Co-authored-by: David Loeffler --- Mathlib.lean | 1 + Mathlib/GroupTheory/Goursat.lean | 178 +++++++++++++++++++++++++++++++ 2 files changed, 179 insertions(+) create mode 100644 Mathlib/GroupTheory/Goursat.lean diff --git a/Mathlib.lean b/Mathlib.lean index d1faa91f26155..624002fb45a2f 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3130,6 +3130,7 @@ import Mathlib.GroupTheory.FreeGroup.Basic import Mathlib.GroupTheory.FreeGroup.IsFreeGroup import Mathlib.GroupTheory.FreeGroup.NielsenSchreier import Mathlib.GroupTheory.FreeGroup.Reduce +import Mathlib.GroupTheory.Goursat import Mathlib.GroupTheory.GroupAction.Basic import Mathlib.GroupTheory.GroupAction.Blocks import Mathlib.GroupTheory.GroupAction.CardCommute diff --git a/Mathlib/GroupTheory/Goursat.lean b/Mathlib/GroupTheory/Goursat.lean new file mode 100644 index 0000000000000..90bc8ec652129 --- /dev/null +++ b/Mathlib/GroupTheory/Goursat.lean @@ -0,0 +1,178 @@ +/- +Copyright (c) 2024 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import Mathlib.Algebra.Group.Graph +import Mathlib.Algebra.Group.Subgroup.Basic +import Mathlib.GroupTheory.QuotientGroup.Defs + +/-! +# Goursat's lemma for subgroups + +This file proves Goursat's lemma for subgroups. + +If `I` is a subgroup of `G × H` which projects fully on both factors, then there exist normal +subgroups `G' ≤ G` and `H' ≤ H` such that `G' × H' ≤ I` and the image of `I` in `G ⧸ G' × H ⧸ H'` is +the graph of an isomorphism `G ⧸ G' ≃ H ⧸ H'`. + +`G'` and `H'` can be explicitly constructed as `Subgroup.goursatFst I` and `Subgroup.goursatSnd I` +respectively. +-/ + +open Function Set + +namespace Subgroup +variable {G H : Type*} [Group G] [Group H] {I : Subgroup (G × H)} + (hI₁ : Surjective (Prod.fst ∘ I.subtype)) (hI₂ : Surjective (Prod.snd ∘ I.subtype)) + +variable (I) in +/-- For `I` a subgroup of `G × H`, `I.goursatFst` is the kernel of the projection map `I → H`, +considered as a subgroup of `G`. + +This is the first subgroup appearing in Goursat's lemma. See `Subgroup.goursat`. -/ +@[to_additive +"For `I` a subgroup of `G × H`, `I.goursatFst` is the kernel of the projection map `I → H`, +considered as a subgroup of `G`. + +This is the first subgroup appearing in Goursat's lemma. See `AddSubgroup.goursat`."] +def goursatFst : Subgroup G := + ((MonoidHom.snd G H).comp I.subtype).ker.map ((MonoidHom.fst G H).comp I.subtype) + +variable (I) in +/-- For `I` a subgroup of `G × H`, `I.goursatSnd` is the kernel of the projection map `I → G`, +considered as a subgroup of `H`. + +This is the second subgroup appearing in Goursat's lemma. See `Subgroup.goursat`. -/ +@[to_additive +"For `I` a subgroup of `G × H`, `I.goursatSnd` is the kernel of the projection map `I → G`, +considered as a subgroup of `H`. + +This is the second subgroup appearing in Goursat's lemma. See `AddSubgroup.goursat`."] +def goursatSnd : Subgroup H := + ((MonoidHom.fst G H).comp I.subtype).ker.map ((MonoidHom.snd G H).comp I.subtype) + +@[to_additive (attr := simp)] +lemma mem_goursatFst {g : G} : g ∈ I.goursatFst ↔ (g, 1) ∈ I := by simp [goursatFst] + +@[to_additive (attr := simp)] +lemma mem_goursatSnd {h : H} : h ∈ I.goursatSnd ↔ (1, h) ∈ I := by simp [goursatSnd] + +include hI₁ in +@[to_additive] lemma normal_goursatFst : I.goursatFst.Normal := .map inferInstance _ hI₁ + +include hI₂ in +@[to_additive] lemma normal_goursatSnd : I.goursatSnd.Normal := .map inferInstance _ hI₂ + +include hI₁ hI₂ in +@[to_additive] +lemma mk_goursatFst_eq_iff_mk_goursatSnd_eq {x y : G × H} (hx : x ∈ I) (hy : y ∈ I) : + (x.1 : G ⧸ I.goursatFst) = y.1 ↔ (x.2 : H ⧸ I.goursatSnd) = y.2 := by + have := normal_goursatFst hI₁ + have := normal_goursatSnd hI₂ + rw [eq_comm] + simp [QuotientGroup.eq_iff_div_mem] + constructor <;> intro h + · simpa [Prod.mul_def, Prod.div_def] using div_mem (mul_mem h hx) hy + · simpa [Prod.mul_def, Prod.div_def] using div_mem (mul_mem h hy) hx + +lemma goursatFst_prod_goursatSnd_le : I.goursatFst.prod I.goursatSnd ≤ I := by + rintro ⟨g, h⟩ ⟨hg, hh⟩ + simpa using mul_mem (mem_goursatFst.1 hg) (mem_goursatSnd.1 hh) + +/-- **Goursat's lemma** for a subgroup of with surjective projections. + +If `I` is a subgroup of `G × H` which projects fully on both factors, then there exist normal +subgroups `M ≤ G` and `N ≤ H` such that `G' × H' ≤ I` and the image of `I` in `G ⧸ M × H ⧸ N` is the +graph of an isomorphism `G ⧸ M ≃ H ⧸ N'`. + +`G'` and `H'` can be explicitly constructed as `I.goursatFst` and `I.goursatSnd` respectively. -/ +@[to_additive +"**Goursat's lemma** for a subgroup of with surjective projections. + +If `I` is a subgroup of `G × H` which projects fully on both factors, then there exist normal +subgroups `M ≤ G` and `N ≤ H` such that `G' × H' ≤ I` and the image of `I` in `G ⧸ M × H ⧸ N` is the +graph of an isomorphism `G ⧸ M ≃ H ⧸ N'`. + +`G'` and `H'` can be explicitly constructed as `I.goursatFst` and `I.goursatSnd` respectively."] +lemma goursat_surjective : + have := normal_goursatFst hI₁ + have := normal_goursatSnd hI₂ + ∃ e : G ⧸ I.goursatFst ≃* H ⧸ I.goursatSnd, + (((QuotientGroup.mk' _).prodMap (QuotientGroup.mk' _)).comp I.subtype).range = + e.toMonoidHom.graph := by + have := normal_goursatFst hI₁ + have := normal_goursatSnd hI₂ + exact (((QuotientGroup.mk' I.goursatFst).prodMap + (QuotientGroup.mk' I.goursatSnd)).comp I.subtype).exists_mulEquiv_range_eq_graph + ((QuotientGroup.mk'_surjective _).comp hI₁) ((QuotientGroup.mk'_surjective _).comp hI₂) + fun ⟨x, hx⟩ ⟨y, hy⟩ ↦ mk_goursatFst_eq_iff_mk_goursatSnd_eq hI₁ hI₂ hx hy + +/-- **Goursat's lemma** for an arbitrary subgroup. + +If `I` is a subgroup of `G × H`, then there exist subgroups `G' ≤ G`, `H' ≤ H` and normal subgroups +`M ≤ G'` and `N ≤ H'` such that `M × N ≤ I` and the image of `I` in `G' ⧸ M × H' ⧸ N` is the graph +of an isomorphism `G ⧸ G' ≃ H ⧸ H'`. -/ +@[to_additive +"**Goursat's lemma** for an arbitrary subgroup. + +If `I` is a subgroup of `G × H`, then there exist subgroups `G' ≤ G`, `H' ≤ H` and normal subgroups +`M ≤ G'` and `N ≤ H'` such that `M × N ≤ I` and the image of `I` in `G' ⧸ M × H' ⧸ N` is the graph +of an isomorphism `G ⧸ G' ≃ H ⧸ H'`."] +lemma goursat : + ∃ (G' : Subgroup G) (H' : Subgroup H) (M : Subgroup G') (N : Subgroup H') (_ : M.Normal) + (_ : N.Normal) (e : G' ⧸ M ≃* H' ⧸ N), + I = (e.toMonoidHom.graph.comap <| (QuotientGroup.mk' M).prodMap (QuotientGroup.mk' N)).map + (G'.subtype.prodMap H'.subtype) := by + let G' := I.map (MonoidHom.fst ..) + let H' := I.map (MonoidHom.snd ..) + let P : I →* G' := (MonoidHom.fst ..).subgroupMap I + let Q : I →* H' := (MonoidHom.snd ..).subgroupMap I + let I' : Subgroup (G' × H') := (P.prod Q).range + have hI₁' : Surjective (Prod.fst ∘ I'.subtype) := by + simp only [← MonoidHom.coe_fst, ← MonoidHom.coe_comp, ← MonoidHom.range_eq_top, + MonoidHom.range_comp, Subgroup.range_subtype, I'] + simp only [← MonoidHom.range_comp, MonoidHom.fst_comp_prod, MonoidHom.range_eq_top] + exact (MonoidHom.fst ..).subgroupMap_surjective I + have hI₂' : Surjective (Prod.snd ∘ I'.subtype) := by + simp only [← MonoidHom.coe_snd, ← MonoidHom.coe_comp, ← MonoidHom.range_eq_top, + MonoidHom.range_comp, Subgroup.range_subtype, I'] + simp only [← MonoidHom.range_comp, MonoidHom.fst_comp_prod, MonoidHom.range_eq_top] + exact (MonoidHom.snd ..).subgroupMap_surjective I + have := normal_goursatFst hI₁' + have := normal_goursatSnd hI₂' + obtain ⟨e, he⟩ := goursat_surjective hI₁' hI₂' + refine ⟨I.map (MonoidHom.fst ..), I.map (MonoidHom.snd ..), + I'.goursatFst, I'.goursatSnd, inferInstance, inferInstance, e, ?_⟩ + rw [← he] + simp only [MonoidHom.range_comp, Subgroup.range_subtype, I'] + rw [comap_map_eq_self] + · ext ⟨g, h⟩ + constructor + · intro hgh + simpa only [mem_map, MonoidHom.mem_range, MonoidHom.prod_apply, Subtype.exists, Prod.exists, + MonoidHom.coe_prodMap, coeSubtype, Prod.mk.injEq, Prod.map_apply, MonoidHom.coe_snd, + exists_eq_right, exists_and_right, exists_eq_right_right, MonoidHom.coe_fst] + using ⟨⟨h, hgh⟩, ⟨g, hgh⟩, g, h, hgh, ⟨rfl, rfl⟩⟩ + · simp only [mem_map, MonoidHom.mem_range, MonoidHom.prod_apply, Subtype.exists, Prod.exists, + MonoidHom.coe_prodMap, coeSubtype, Prod.mk.injEq, Prod.map_apply, MonoidHom.coe_snd, + exists_eq_right, exists_and_right, exists_eq_right_right, MonoidHom.coe_fst, + forall_exists_index, and_imp] + rintro h₁ hgh₁ g₁ hg₁h g₂ h₂ hg₂h₂ hP hQ + simp only [Subtype.ext_iff] at hP hQ + rwa [← hP, ← hQ] + · rintro ⟨⟨g, _⟩, ⟨h, _⟩⟩ hgh + simp only [MonoidHom.prodMap, MonoidHom.mem_ker, MonoidHom.prod_apply, MonoidHom.coe_comp, + QuotientGroup.coe_mk', MonoidHom.coe_fst, comp_apply, MonoidHom.coe_snd, Prod.mk_eq_one, + QuotientGroup.eq_one_iff, mem_goursatFst, MonoidHom.mem_range, Prod.mk.injEq, Subtype.exists, + Prod.exists, mem_goursatSnd] at hgh + rcases hgh with ⟨⟨g₁, h₁, hg₁h₁, hPQ₁⟩, ⟨g₂, h₂, hg₂h₂, hPQ₂⟩⟩ + simp only [Subtype.ext_iff] at hPQ₁ hPQ₂ + rcases hPQ₁ with ⟨rfl, rfl⟩ + rcases hPQ₂ with ⟨rfl, rfl⟩ + simp only [MonoidHom.mem_range, MonoidHom.prod_apply, Prod.mk.injEq, Subtype.exists, + Prod.exists] + refine ⟨g₁, h₂, ?_, rfl, rfl⟩ + simpa only [OneMemClass.coe_one, Prod.mk_mul_mk, mul_one, one_mul] using mul_mem hg₁h₁ hg₂h₂ + +end Subgroup From 2b0dc338583b56ac2dc2eaf8caf8f4fa71892317 Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Thu, 28 Nov 2024 16:12:26 +0000 Subject: [PATCH 390/829] feat: drop finite-dimensionality assumption from `PerfectPairing.restrictScalarsField` (#18940) --- Mathlib.lean | 1 + Mathlib/Data/Matrix/Mul.lean | 12 +++ Mathlib/LinearAlgebra/Dual.lean | 2 +- Mathlib/LinearAlgebra/Matrix/BaseChange.lean | 60 +++++++++++++ Mathlib/LinearAlgebra/Matrix/Basis.lean | 6 ++ Mathlib/LinearAlgebra/PerfectPairing.lean | 93 ++++++++++++++++++-- docs/references.bib | 2 +- 7 files changed, 165 insertions(+), 11 deletions(-) create mode 100644 Mathlib/LinearAlgebra/Matrix/BaseChange.lean diff --git a/Mathlib.lean b/Mathlib.lean index 624002fb45a2f..b270a2e2785a0 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3349,6 +3349,7 @@ import Mathlib.LinearAlgebra.LinearIndependent import Mathlib.LinearAlgebra.LinearPMap import Mathlib.LinearAlgebra.Matrix.AbsoluteValue import Mathlib.LinearAlgebra.Matrix.Adjugate +import Mathlib.LinearAlgebra.Matrix.BaseChange import Mathlib.LinearAlgebra.Matrix.Basis import Mathlib.LinearAlgebra.Matrix.BilinearForm import Mathlib.LinearAlgebra.Matrix.Block diff --git a/Mathlib/Data/Matrix/Mul.lean b/Mathlib/Data/Matrix/Mul.lean index e0d778836f85b..16985ffe09158 100644 --- a/Mathlib/Data/Matrix/Mul.lean +++ b/Mathlib/Data/Matrix/Mul.lean @@ -880,6 +880,18 @@ theorem submatrix_mulVec_equiv [Fintype n] [Fintype o] [NonUnitalNonAssocSemirin M.submatrix e₁ e₂ *ᵥ v = (M *ᵥ (v ∘ e₂.symm)) ∘ e₁ := funext fun _ => Eq.symm (dotProduct_comp_equiv_symm _ _ _) +@[simp] +theorem submatrix_id_mul_left [Fintype n] [Fintype o] [Mul α] [AddCommMonoid α] {p : Type*} + (M : Matrix m n α) (N : Matrix o p α) (e₁ : l → m) (e₂ : n ≃ o) : + M.submatrix e₁ id * N.submatrix e₂ id = M.submatrix e₁ e₂.symm * N := by + ext; simp [mul_apply, ← e₂.bijective.sum_comp] + +@[simp] +theorem submatrix_id_mul_right [Fintype n] [Fintype o] [Mul α] [AddCommMonoid α] {p : Type*} + (M : Matrix m n α) (N : Matrix o p α) (e₁ : l → p) (e₂ : o ≃ n) : + M.submatrix id e₂ * N.submatrix id e₁ = M * N.submatrix e₂.symm e₁ := by + ext; simp [mul_apply, ← e₂.bijective.sum_comp] + theorem submatrix_vecMul_equiv [Fintype l] [Fintype m] [NonUnitalNonAssocSemiring α] (M : Matrix m n α) (v : l → α) (e₁ : l ≃ m) (e₂ : o → n) : v ᵥ* M.submatrix e₁ e₂ = ((v ∘ e₁.symm) ᵥ* M) ∘ e₂ := diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index a203905b0a8a2..3870b51c3d7f0 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -418,7 +418,7 @@ theorem linearCombination_dualBasis (f : ι →₀ R) (i : ι) : @[deprecated (since := "2024-08-29")] alias total_dualBasis := linearCombination_dualBasis -theorem dualBasis_repr (l : Dual R M) (i : ι) : b.dualBasis.repr l i = l (b i) := by +@[simp] theorem dualBasis_repr (l : Dual R M) (i : ι) : b.dualBasis.repr l i = l (b i) := by rw [← linearCombination_dualBasis b, Basis.linearCombination_repr b.dualBasis l] theorem dualBasis_apply (i : ι) (m : M) : b.dualBasis i m = b.repr m i := diff --git a/Mathlib/LinearAlgebra/Matrix/BaseChange.lean b/Mathlib/LinearAlgebra/Matrix/BaseChange.lean new file mode 100644 index 0000000000000..7261b7f111eab --- /dev/null +++ b/Mathlib/LinearAlgebra/Matrix/BaseChange.lean @@ -0,0 +1,60 @@ +/- +Copyright (c) 2024 Oliver Nash. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Oliver Nash +-/ +import Mathlib.LinearAlgebra.Matrix.NonsingularInverse +import Mathlib.Algebra.Field.Subfield.Defs + +/-! +# Matrices and base change + +This file is a home for results about base change for matrices. + +## Main results: + * `Matrix.mem_subfield_of_mul_eq_one_of_mem_subfield_right`: if an invertible matrix over `L` takes + values in subfield `K ⊆ L`, then so does its (right) inverse. + * `Matrix.mem_subfield_of_mul_eq_one_of_mem_subfield_left`: if an invertible matrix over `L` takes + values in subfield `K ⊆ L`, then so does its (left) inverse. + +-/ + +namespace Matrix + +open scoped Matrix + +variable {m n L : Type*} [Fintype m] [Fintype n] [DecidableEq m] [Field L] + (e : m ≃ n) (K : Subfield L) {A : Matrix m n L} {B : Matrix n m L} (hAB : A * B = 1) + +include e hAB + +lemma mem_subfield_of_mul_eq_one_of_mem_subfield_right + (h_mem : ∀ i j, A i j ∈ K) (i : n) (j : m) : + B i j ∈ K := by + let A' : Matrix m m K := of fun i j ↦ ⟨A.submatrix id e i j, h_mem i (e j)⟩ + have hA' : A'.map K.subtype = A.submatrix id e := rfl + have hA : IsUnit A' := by + have h_unit : IsUnit (A.submatrix id e) := + isUnit_of_right_inverse (B := B.submatrix e id) (by simpa) + have h_det : (A.submatrix id e).det = K.subtype A'.det := by + simp [A', K.subtype.map_det, map, submatrix] + simpa [isUnit_iff_isUnit_det, h_det] using h_unit + obtain ⟨B', hB⟩ := exists_right_inverse_iff_isUnit.mpr hA + suffices (B'.submatrix e.symm id).map K.subtype = B by simp [← this] + replace hB : A * (B'.submatrix e.symm id).map K.subtype = 1 := by + replace hB := congr_arg (fun C ↦ C.map K.subtype) hB + simp_rw [map_mul] at hB + rw [hA', ← e.symm_symm, ← submatrix_id_mul_left] at hB + simpa using hB + classical + simpa [← Matrix.mul_assoc, (mul_eq_one_comm_of_equiv e).mp hAB] using congr_arg (B * ·) hB + +lemma mem_subfield_of_mul_eq_one_of_mem_subfield_left + (h_mem : ∀ i j, B i j ∈ K) (i : m) (j : n) : + A i j ∈ K := by + replace hAB : Bᵀ * Aᵀ = 1 := by simpa using congr_arg transpose hAB + rw [← A.transpose_apply] + simp_rw [← B.transpose_apply] at h_mem + exact mem_subfield_of_mul_eq_one_of_mem_subfield_right e K hAB (fun i j ↦ h_mem j i) j i + +end Matrix diff --git a/Mathlib/LinearAlgebra/Matrix/Basis.lean b/Mathlib/LinearAlgebra/Matrix/Basis.lean index 1fb45927bb7b4..f06b9e9d99955 100644 --- a/Mathlib/LinearAlgebra/Matrix/Basis.lean +++ b/Mathlib/LinearAlgebra/Matrix/Basis.lean @@ -230,6 +230,12 @@ theorem Basis.toMatrix_reindex' [DecidableEq ι] [DecidableEq ι'] (b : Basis ι Matrix.reindex_apply, Matrix.submatrix_apply, Function.comp_apply, e.apply_symm_apply, Finsupp.mapDomain_equiv_apply] +@[simp] +lemma Basis.toMatrix_mulVec_repr (m : M) : + b'.toMatrix b *ᵥ b.repr m = b'.repr m := by + classical + simp [← LinearMap.toMatrix_id_eq_basis_toMatrix, LinearMap.toMatrix_mulVec_repr] + end Fintype /-- A generalization of `Basis.toMatrix_self`, in the opposite direction. -/ diff --git a/Mathlib/LinearAlgebra/PerfectPairing.lean b/Mathlib/LinearAlgebra/PerfectPairing.lean index 0b827f3a4969c..ca394d510fa52 100644 --- a/Mathlib/LinearAlgebra/PerfectPairing.lean +++ b/Mathlib/LinearAlgebra/PerfectPairing.lean @@ -4,6 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash -/ import Mathlib.LinearAlgebra.Dual +import Mathlib.LinearAlgebra.Matrix.Basis +import Mathlib.LinearAlgebra.Matrix.BaseChange +import Mathlib.LinearAlgebra.FreeModule.Finite.Matrix /-! # Perfect pairings of modules @@ -21,6 +24,7 @@ to connect 1 and 2. * `PerfectPairing` * `PerfectPairing.flip` + * `PerfectPairing.dual` * `PerfectPairing.toDualLeft` * `PerfectPairing.toDualRight` * `PerfectPairing.restrict` @@ -161,6 +165,11 @@ include p in theorem reflexive_right : IsReflexive R N := p.flip.reflexive_left +include p in +theorem finrank_eq [Module.Finite R M] [Module.Free R M] : + finrank R M = finrank R N := + ((Module.Free.chooseBasis R M).toDualEquiv.trans p.toDualRight.symm).finrank_eq + private lemma restrict_aux {M' N' : Type*} [AddCommGroup M'] [Module R M'] [AddCommGroup N'] [Module R N'] (i : M' →ₗ[R] M) (j : N' →ₗ[R] N) @@ -206,6 +215,61 @@ section RestrictScalars open Submodule (span) +/-- If a perfect pairing over a field `L` takes values in a subfield `K` along two `K`-subspaces +whose `L` span is full, then these subspaces induce a `K`-structure in the sense of +[*Algebra I*, Bourbaki : Chapter II, §8.1 Definition 1][bourbaki1989]. -/ +lemma exists_basis_basis_of_span_eq_top_of_mem_algebraMap + {K L : Type*} [Field K] [Field L] [Algebra K L] + [Module L M] [Module L N] [Module K M] [Module K N] [IsScalarTower K L M] + (p : PerfectPairing L M N) + (M' : Submodule K M) (N' : Submodule K N) + (hM : span L (M' : Set M) = ⊤) + (hN : span L (N' : Set N) = ⊤) + (hp : ∀ᵉ (x ∈ M') (y ∈ N'), p x y ∈ (algebraMap K L).range) : + ∃ (n : ℕ) (b : Basis (Fin n) L M) (b' : Basis (Fin n) K M'), ∀ i, b i = b' i := by + classical + have : IsReflexive L M := p.reflexive_left + have : IsReflexive L N := p.reflexive_right + obtain ⟨v, hv₁, hv₂, hv₃⟩ := exists_linearIndependent L (M' : Set M) + rw [hM] at hv₂ + let b : Basis _ L M := Basis.mk hv₃ <| by rw [← hv₂, Subtype.range_coe_subtype, Set.setOf_mem_eq] + have : Fintype v := Set.Finite.fintype <| Module.Finite.finite_basis b + set v' : v → M' := fun i ↦ ⟨i, hv₁ (Subtype.coe_prop i)⟩ + have hv' : LinearIndependent K v' := by + replace hv₃ := hv₃.restrict_scalars (R := K) <| by + simp_rw [← Algebra.algebraMap_eq_smul_one] + exact NoZeroSMulDivisors.algebraMap_injective K L + rw [show ((↑) : v → M) = M'.subtype ∘ v' from rfl] at hv₃ + exact hv₃.of_comp + suffices span K (Set.range v') = ⊤ by + let e := (Module.Finite.finite_basis b).equivFin + let b' : Basis _ K M' := Basis.mk hv' (by rw [this]) + exact ⟨_, b.reindex e, b'.reindex e, fun i ↦ by simp [b, b']⟩ + suffices span K v = M' by + apply Submodule.map_injective_of_injective M'.injective_subtype + rw [Submodule.map_span, ← Set.image_univ, Set.image_image] + simpa + refine le_antisymm (Submodule.span_le.mpr hv₁) fun m hm ↦ ?_ + obtain ⟨w, hw₁, hw₂, hw₃⟩ := exists_linearIndependent L (N' : Set N) + rw [hN] at hw₂ + let bN : Basis _ L N := Basis.mk hw₃ <| by rw [← hw₂, Subtype.range_coe_subtype, Set.setOf_mem_eq] + have : Fintype w := Set.Finite.fintype <| Module.Finite.finite_basis bN + have e : v ≃ w := Fintype.equivOfCardEq <| by rw [← Module.finrank_eq_card_basis b, + ← Module.finrank_eq_card_basis bN, p.finrank_eq] + let bM := bN.dualBasis.map p.toDualLeft.symm + have hbM (j : w) (x : M) (hx : x ∈ M') : bM.repr x j = p x (j : N) := by simp [bM, bN] + have hj (j : w) : bM.repr m j ∈ (algebraMap K L).range := (hbM _ _ hm) ▸ hp m hm j (hw₁ j.2) + replace hp (i : w) (j : v) : + (bN.dualBasis.map p.toDualLeft.symm).toMatrix b i j ∈ (algebraMap K L).fieldRange := by + simp only [Basis.toMatrix, Basis.map_repr, LinearEquiv.symm_symm, LinearEquiv.trans_apply, + toDualLeft_apply, Basis.dualBasis_repr] + exact hp (b j) (by simpa [b] using hv₁ j.2) (bN i) (by simpa [bN] using hw₁ i.2) + have hA (i j) : b.toMatrix bM i j ∈ (algebraMap K L).range := + Matrix.mem_subfield_of_mul_eq_one_of_mem_subfield_left e _ (by simp) hp i j + have h_span : span K v = span K (Set.range b) := by simp [b] + rw [h_span, Basis.mem_span_iff_repr_mem, ← Basis.toMatrix_mulVec_repr bM b m] + exact fun i ↦ Subring.sum_mem _ fun j _ ↦ Subring.mul_mem _ (hA i j) (hj j) + variable {S : Type*} [CommRing S] [Algebra S R] [Module S M] [Module S N] [IsScalarTower S R M] [IsScalarTower S R N] [NoZeroSMulDivisors S R] [Nontrivial R] @@ -272,18 +336,23 @@ def restrictScalars p.flip.restrictScalarsAux_surjective j i h₂ (fun m n ↦ hp n m)⟩} /-- Restriction of scalars for a perfect pairing taking values in a subfield. -/ -def restrictScalarsField {K : Type*} [Field K] [Algebra K R] - [Module K M] [Module K N] [IsScalarTower K R M] [IsScalarTower K R N] - [Module K M'] [Module K N'] [FiniteDimensional K M'] +def restrictScalarsField {K L : Type*} [Field K] [Field L] [Algebra K L] + [Module L M] [Module L N] [Module K M] [Module K N] [IsScalarTower K L M] [IsScalarTower K L N] + [Module K M'] [Module K N'] (i : M' →ₗ[K] M) (j : N' →ₗ[K] N) (hi : Injective i) (hj : Injective j) - (hM : span R (LinearMap.range i : Set M) = ⊤) - (hN : span R (LinearMap.range j : Set N) = ⊤) - (hp : ∀ m n, p (i m) (j n) ∈ (algebraMap K R).range) : - PerfectPairing K M' N' := - PerfectPairing.mkOfInjective _ - (p.restrictScalarsAux_injective i j hi hN hp) + (hM : span L (LinearMap.range i : Set M) = ⊤) + (hN : span L (LinearMap.range j : Set N) = ⊤) + (p : PerfectPairing L M N) + (hp : ∀ m n, p (i m) (j n) ∈ (algebraMap K L).range) : + PerfectPairing K M' N' := by + suffices FiniteDimensional K M' from mkOfInjective _ (p.restrictScalarsAux_injective i j hi hN hp) (p.flip.restrictScalarsAux_injective j i hj hM (fun m n ↦ hp n m)) + obtain ⟨n, -, b', -⟩ := p.exists_basis_basis_of_span_eq_top_of_mem_algebraMap _ _ hM hN <| by + rintro - ⟨m, rfl⟩ - ⟨n, rfl⟩ + exact hp m n + have : FiniteDimensional K (LinearMap.range i) := FiniteDimensional.of_fintype_basis b' + exact Finite.equiv (LinearEquiv.ofInjective i hi).symm end RestrictScalars @@ -341,6 +410,12 @@ def toPerfectPairing : PerfectPairing R N M where end LinearEquiv +/-- A perfect pairing induces a perfect pairing between dual spaces. -/ +def PerfectPairing.dual (p : PerfectPairing R M N) : + PerfectPairing R (Dual R M) (Dual R N) := + let _i := p.reflexive_right + (p.toDualRight.symm.trans (evalEquiv R N)).toPerfectPairing + namespace Submodule open LinearEquiv diff --git a/docs/references.bib b/docs/references.bib index 0e5ce6b9bd094..976afa9ae6d69 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -607,7 +607,7 @@ @Book{ bourbaki1987 @Book{ bourbaki1989, author = {Bourbaki, N.}, - title = {Algebra. {II}. {C}hapters 1–3}, + title = {Algebra. {I}. {C}hapters 1–3}, series = {Elements of Mathematics}, note = {Translated from the French}, year = {1998}, From 0be336ffe18f3501a96f8248be4bd6778eec8a38 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 28 Nov 2024 16:12:28 +0000 Subject: [PATCH 391/829] feat: `MulLeftMono` implies `PosMulMono` (#19248) ... and similarly for the other typeclasses From GrowthInGroups (LeanCamCombi) Co-authored-by: Andrew Yang --- .../Order/GroupWithZero/Unbundled.lean | 24 +++++++++++++++++++ 1 file changed, 24 insertions(+) diff --git a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean index 488d67de45348..3a29eb4977c1c 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean @@ -193,6 +193,30 @@ instance MulPosReflectLT.to_contravariantClass_pos_mul_lt [MulPosReflectLT α] : ContravariantClass α>0 α (fun x y => y * x) (· < ·) := ⟨fun a _ _ bc => @ContravariantClass.elim α≥0 α (fun x y => y * x) (· < ·) _ ⟨_, a.2.le⟩ _ _ bc⟩ +instance (priority := 100) MulLeftMono.toPosMulMono [MulLeftMono α] : + PosMulMono α where elim _ _ := ‹MulLeftMono α›.elim _ + +instance (priority := 100) MulRightMono.toMulPosMono [MulRightMono α] : + MulPosMono α where elim _ _ := ‹MulRightMono α›.elim _ + +instance (priority := 100) MulLeftStrictMono.toPosMulStrictMono [MulLeftStrictMono α] : + PosMulStrictMono α where elim _ _ := ‹MulLeftStrictMono α›.elim _ + +instance (priority := 100) MulRightStrictMono.toMulPosStrictMono [MulRightStrictMono α] : + MulPosStrictMono α where elim _ _ := ‹MulRightStrictMono α›.elim _ + +instance (priority := 100) MulLeftMono.toPosMulReflectLT [MulLeftReflectLT α] : + PosMulReflectLT α where elim _ _ := ‹MulLeftReflectLT α›.elim _ + +instance (priority := 100) MulRightMono.toMulPosReflectLT [MulRightReflectLT α] : + MulPosReflectLT α where elim _ _ := ‹MulRightReflectLT α›.elim _ + +instance (priority := 100) MulLeftStrictMono.toPosMulReflectLE [MulLeftReflectLE α] : + PosMulReflectLE α where elim _ _ := ‹MulLeftReflectLE α›.elim _ + +instance (priority := 100) MulRightStrictMono.toMulPosReflectLE [MulRightReflectLE α] : + MulPosReflectLE α where elim _ _ := ‹MulRightReflectLE α›.elim _ + @[gcongr] theorem mul_le_mul_of_nonneg_left [PosMulMono α] (h : b ≤ c) (a0 : 0 ≤ a) : a * b ≤ a * c := @CovariantClass.elim α≥0 α (fun x y => x * y) (· ≤ ·) _ ⟨a, a0⟩ _ _ h From 37dfa236e97da1f388140bf57e71abe9fcb39166 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 28 Nov 2024 16:12:29 +0000 Subject: [PATCH 392/829] chore(Analysis/Multilinear): avoid explicit `variable`s (#19445) Also change implicitness of some arguments, including the function argument of `ContinuousMultilinearMap.opNorm_le_bound`, expand some docstrings, and reorder arguments of `ContinuousMultilinearMap.le_of_opNorm_le` --- Mathlib/Analysis/Analytic/ChangeOrigin.lean | 2 +- Mathlib/Analysis/Analytic/Composition.lean | 2 +- .../Calculus/ContDiff/FaaDiBruno.lean | 2 +- .../Fourier/FourierTransformDeriv.lean | 2 +- .../Normed/Operator/BoundedLinearMaps.lean | 34 +-- .../NormedSpace/Multilinear/Basic.lean | 241 +++++++++--------- .../NormedSpace/Multilinear/Curry.lean | 4 +- .../PiTensorProduct/InjectiveSeminorm.lean | 4 +- 8 files changed, 138 insertions(+), 153 deletions(-) diff --git a/Mathlib/Analysis/Analytic/ChangeOrigin.lean b/Mathlib/Analysis/Analytic/ChangeOrigin.lean index f07c1dab4aa20..45dce2a33326c 100644 --- a/Mathlib/Analysis/Analytic/ChangeOrigin.lean +++ b/Mathlib/Analysis/Analytic/ChangeOrigin.lean @@ -106,7 +106,7 @@ theorem nnnorm_changeOriginSeries_apply_le_tsum (k l : ℕ) (x : E) : ‖p.changeOriginSeries k l fun _ => x‖₊ ≤ ∑' _ : { s : Finset (Fin (k + l)) // s.card = l }, ‖p (k + l)‖₊ * ‖x‖₊ ^ l := by rw [NNReal.tsum_mul_right, ← Fin.prod_const] - exact (p.changeOriginSeries k l).le_of_opNNNorm_le _ (p.nnnorm_changeOriginSeries_le_tsum _ _) + exact (p.changeOriginSeries k l).le_of_opNNNorm_le (p.nnnorm_changeOriginSeries_le_tsum _ _) _ /-- Changing the origin of a formal multilinear series `p`, so that `p.sum (x+y) = (p.changeOrigin x).sum y` when this makes sense. diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index ef4d2f8bb4054..472d06dd79024 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -309,7 +309,7 @@ the norms of the relevant bits of `q` and `p`. -/ theorem compAlongComposition_norm {n : ℕ} (q : FormalMultilinearSeries 𝕜 F G) (p : FormalMultilinearSeries 𝕜 E F) (c : Composition n) : ‖q.compAlongComposition p c‖ ≤ ‖q c.length‖ * ∏ i, ‖p (c.blocksFun i)‖ := - ContinuousMultilinearMap.opNorm_le_bound _ (by positivity) (compAlongComposition_bound _ _ _) + ContinuousMultilinearMap.opNorm_le_bound (by positivity) (compAlongComposition_bound _ _ _) theorem compAlongComposition_nnnorm {n : ℕ} (q : FormalMultilinearSeries 𝕜 F G) (p : FormalMultilinearSeries 𝕜 E F) (c : Composition n) : diff --git a/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean b/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean index 582c9ca4fec35..a022b455054dd 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean @@ -792,7 +792,7 @@ noncomputable def compAlongOrderedFinpartitionL : refine MultilinearMap.mkContinuousLinear c.compAlongOrderedFinpartitionₗ 1 (fun f p ↦ ?_) simp only [one_mul] change ‖c.compAlongOrderedFinpartition f p‖ ≤ _ - apply ContinuousMultilinearMap.opNorm_le_bound _ (by positivity) (fun v ↦ ?_) + apply ContinuousMultilinearMap.opNorm_le_bound (by positivity) (fun v ↦ ?_) simp only [compAlongOrderFinpartition_apply] apply (f.le_opNorm _).trans rw [mul_assoc, ← c.prod_sigma_eq_prod, ← Finset.prod_mul_distrib] diff --git a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean index dcf554692ab17..d5e16eb0996f7 100644 --- a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean +++ b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean @@ -322,7 +322,7 @@ lemma _root_.ContDiff.fourierPowSMulRight lemma norm_fourierPowSMulRight_le (f : V → E) (v : V) (n : ℕ) : ‖fourierPowSMulRight L f v n‖ ≤ (2 * π * ‖L‖) ^ n * ‖v‖ ^ n * ‖f v‖ := by - apply ContinuousMultilinearMap.opNorm_le_bound _ (by positivity) (fun m ↦ ?_) + apply ContinuousMultilinearMap.opNorm_le_bound (by positivity) (fun m ↦ ?_) calc ‖fourierPowSMulRight L f v n m‖ = (2 * π) ^ n * ((∏ x : Fin n, |(L v) (m x)|) * ‖f v‖) := by diff --git a/Mathlib/Analysis/Normed/Operator/BoundedLinearMaps.lean b/Mathlib/Analysis/Normed/Operator/BoundedLinearMaps.lean index b94e46d1c8ddf..33d288c85a3c2 100644 --- a/Mathlib/Analysis/Normed/Operator/BoundedLinearMaps.lean +++ b/Mathlib/Analysis/Normed/Operator/BoundedLinearMaps.lean @@ -189,40 +189,16 @@ operation. -/ theorem isBoundedLinearMap_prod_multilinear {E : ι → Type*} [∀ i, SeminormedAddCommGroup (E i)] [∀ i, NormedSpace 𝕜 (E i)] : IsBoundedLinearMap 𝕜 fun p : ContinuousMultilinearMap 𝕜 E F × ContinuousMultilinearMap 𝕜 E G => - p.1.prod p.2 where - map_add p₁ p₂ := by ext : 1; rfl - map_smul c p := by ext : 1; rfl - bound := by - refine ⟨1, zero_lt_one, fun p ↦ ?_⟩ - rw [one_mul] - apply ContinuousMultilinearMap.opNorm_le_bound _ (norm_nonneg _) _ - intro m - rw [ContinuousMultilinearMap.prod_apply, norm_prod_le_iff] - constructor - · exact (p.1.le_opNorm m).trans (mul_le_mul_of_nonneg_right (norm_fst_le p) <| by positivity) - · exact (p.2.le_opNorm m).trans (mul_le_mul_of_nonneg_right (norm_snd_le p) <| by positivity) + p.1.prod p.2 := + (ContinuousMultilinearMap.prodL ..).toContinuousLinearEquiv + |>.toContinuousLinearMap.isBoundedLinearMap /-- Given a fixed continuous linear map `g`, associating to a continuous multilinear map `f` the continuous multilinear map `f (g m₁, ..., g mₙ)` is a bounded linear operation. -/ theorem isBoundedLinearMap_continuousMultilinearMap_comp_linear (g : G →L[𝕜] E) : IsBoundedLinearMap 𝕜 fun f : ContinuousMultilinearMap 𝕜 (fun _ : ι => E) F => - f.compContinuousLinearMap fun _ => g := by - refine - IsLinearMap.with_bound - ⟨fun f₁ f₂ => by ext; rfl, - fun c f => by ext; rfl⟩ - (‖g‖ ^ Fintype.card ι) fun f => ?_ - apply ContinuousMultilinearMap.opNorm_le_bound _ _ _ - · apply_rules [mul_nonneg, pow_nonneg, norm_nonneg] - intro m - calc - ‖f (g ∘ m)‖ ≤ ‖f‖ * ∏ i, ‖g (m i)‖ := f.le_opNorm _ - _ ≤ ‖f‖ * ∏ i, ‖g‖ * ‖m i‖ := by - apply mul_le_mul_of_nonneg_left _ (norm_nonneg _) - exact Finset.prod_le_prod (fun i _ => norm_nonneg _) fun i _ => g.le_opNorm _ - _ = ‖g‖ ^ Fintype.card ι * ‖f‖ * ∏ i, ‖m i‖ := by - simp only [Finset.prod_mul_distrib, Finset.prod_const, Finset.card_univ] - ring + f.compContinuousLinearMap fun _ => g := + (ContinuousMultilinearMap.compContinuousLinearMapL fun _ ↦ g).isBoundedLinearMap end diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean index 95b2d5b3b4aaa..9ebb2a8c9c154 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean @@ -34,8 +34,7 @@ Let `f` be a continuous multilinear map in finitely many variables. ## Implementation notes We mostly follow the API (and the proofs) of `OperatorNorm.lean`, with the additional complexity -that we should deal with multilinear maps in several variables. The currying/uncurrying -constructions are based on those in `Multilinear.lean`. +that we should deal with multilinear maps in several variables. From the mathematical point of view, all the results follow from the results on operator norm in one variable, by applying them to one variable after the other through currying. However, this @@ -51,14 +50,6 @@ noncomputable section open scoped NNReal Topology Uniformity open Finset Metric Function Filter -/- -Porting note: These lines are not required in Mathlib4. -```lean -attribute [local instance 1001] - AddCommGroup.toAddCommMonoid NormedAddCommGroup.toAddCommGroup NormedSpace.toModule' -``` --/ - /-! ### Type variables @@ -97,21 +88,20 @@ protected alias ContinuousMultilinearMap.continuous_eval := continuous_eval namespace ContinuousLinearMap variable {G : Type*} [AddCommGroup G] [TopologicalSpace G] [Module 𝕜 G] [ContinuousConstSMul 𝕜 F] - (f : G →L[𝕜] ContinuousMultilinearMap 𝕜 E F) -lemma continuous_uncurry_of_multilinear : +lemma continuous_uncurry_of_multilinear (f : G →L[𝕜] ContinuousMultilinearMap 𝕜 E F) : Continuous (fun (p : G × (Π i, E i)) ↦ f p.1 p.2) := by fun_prop -lemma continuousOn_uncurry_of_multilinear {s} : +lemma continuousOn_uncurry_of_multilinear (f : G →L[𝕜] ContinuousMultilinearMap 𝕜 E F) {s} : ContinuousOn (fun (p : G × (Π i, E i)) ↦ f p.1 p.2) s := f.continuous_uncurry_of_multilinear.continuousOn -lemma continuousAt_uncurry_of_multilinear {x} : +lemma continuousAt_uncurry_of_multilinear (f : G →L[𝕜] ContinuousMultilinearMap 𝕜 E F) {x} : ContinuousAt (fun (p : G × (Π i, E i)) ↦ f p.1 p.2) x := f.continuous_uncurry_of_multilinear.continuousAt -lemma continuousWithinAt_uncurry_of_multilinear {s x} : +lemma continuousWithinAt_uncurry_of_multilinear (f : G →L[𝕜] ContinuousMultilinearMap 𝕜 E F) {s x} : ContinuousWithinAt (fun (p : G × (Π i, E i)) ↦ f p.1 p.2) s x := f.continuous_uncurry_of_multilinear.continuousWithinAt @@ -136,16 +126,15 @@ both directions. Along the way, we prove useful bounds on the difference `‖f m namespace MultilinearMap -variable (f : MultilinearMap 𝕜 E G) - -/-- If `f` is a continuous multilinear map in finitely many variables on `E` and `m` is an element -of `∀ i, E i` such that one of the `m i` has norm `0`, then `f m` has norm `0`. +/-- If `f` is a continuous multilinear map on `E` +and `m` is an element of `∀ i, E i` such that one of the `m i` has norm `0`, +then `f m` has norm `0`. Note that we cannot drop the continuity assumption because `f (m : Unit → E) = f (m ())`, where the domain has zero norm and the codomain has a nonzero norm does not satisfy this condition. -/ -lemma norm_map_coord_zero (hf : Continuous f) {m : ∀ i, E i} {i : ι} (hi : ‖m i‖ = 0) : - ‖f m‖ = 0 := by +lemma norm_map_coord_zero (f : MultilinearMap 𝕜 E G) (hf : Continuous f) + {m : ∀ i, E i} {i : ι} (hi : ‖m i‖ = 0) : ‖f m‖ = 0 := by classical rw [← inseparable_zero_iff_norm] at hi ⊢ have : Inseparable (update m i 0) m := inseparable_pi.2 <| @@ -154,7 +143,17 @@ lemma norm_map_coord_zero (hf : Continuous f) {m : ∀ i, E i} {i : ι} (hi : variable [Fintype ι] -theorem bound_of_shell_of_norm_map_coord_zero (hf₀ : ∀ {m i}, ‖m i‖ = 0 → ‖f m‖ = 0) +/-- If a multilinear map in finitely many variables on seminormed spaces +sends vectors with a component of norm zero to vectors of norm zero +and satisfies the inequality `‖f m‖ ≤ C * ∏ i, ‖m i‖` on a shell `ε i / ‖c i‖ < ‖m i‖ < ε i` +for some positive numbers `ε i` and elements `c i : 𝕜`, `1 < ‖c i‖`, +then it satisfies this inequality for all `m`. + +The first assumption is automatically satisfied on normed spaces, see `bound_of_shell` below. +For seminormed spaces, it follows from continuity of `f`, see next lemma, +see `bound_of_shell_of_continuous` below. -/ +theorem bound_of_shell_of_norm_map_coord_zero (f : MultilinearMap 𝕜 E G) + (hf₀ : ∀ {m i}, ‖m i‖ = 0 → ‖f m‖ = 0) {ε : ι → ℝ} {C : ℝ} (hε : ∀ i, 0 < ε i) {c : ι → 𝕜} (hc : ∀ i, 1 < ‖c i‖) (hf : ∀ m : ∀ i, E i, (∀ i, ε i / ‖c i‖ ≤ ‖m i‖) → (∀ i, ‖m i‖ < ε i) → ‖f m‖ ≤ C * ∏ i, ‖m i‖) (m : ∀ i, E i) : ‖f m‖ ≤ C * ∏ i, ‖m i‖ := by @@ -169,7 +168,7 @@ theorem bound_of_shell_of_norm_map_coord_zero (hf₀ : ∀ {m i}, ‖m i‖ = 0 /-- If a continuous multilinear map in finitely many variables on normed spaces satisfies the inequality `‖f m‖ ≤ C * ∏ i, ‖m i‖` on a shell `ε i / ‖c i‖ < ‖m i‖ < ε i` for some positive numbers `ε i` and elements `c i : 𝕜`, `1 < ‖c i‖`, then it satisfies this inequality for all `m`. -/ -theorem bound_of_shell_of_continuous (hfc : Continuous f) +theorem bound_of_shell_of_continuous (f : MultilinearMap 𝕜 E G) (hfc : Continuous f) {ε : ι → ℝ} {C : ℝ} (hε : ∀ i, 0 < ε i) {c : ι → 𝕜} (hc : ∀ i, 1 < ‖c i‖) (hf : ∀ m : ∀ i, E i, (∀ i, ε i / ‖c i‖ ≤ ‖m i‖) → (∀ i, ‖m i‖ < ε i) → ‖f m‖ ≤ C * ∏ i, ‖m i‖) (m : ∀ i, E i) : ‖f m‖ ≤ C * ∏ i, ‖m i‖ := @@ -178,7 +177,7 @@ theorem bound_of_shell_of_continuous (hfc : Continuous f) /-- If a multilinear map in finitely many variables on normed spaces is continuous, then it satisfies the inequality `‖f m‖ ≤ C * ∏ i, ‖m i‖`, for some `C` which can be chosen to be positive. -/ -theorem exists_bound_of_continuous (hf : Continuous f) : +theorem exists_bound_of_continuous (f : MultilinearMap 𝕜 E G) (hf : Continuous f) : ∃ C : ℝ, 0 < C ∧ ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖ := by cases isEmpty_or_nonempty ι · refine ⟨‖f 0‖ + 1, add_pos_of_nonneg_of_pos (norm_nonneg _) zero_lt_one, fun m => ?_⟩ @@ -195,13 +194,15 @@ theorem exists_bound_of_continuous (hf : Continuous f) : rw [← div_le_iff₀' this, one_div, ← inv_pow, inv_div, Fintype.card, ← prod_const] exact prod_le_prod (fun _ _ => div_nonneg ε0.le (norm_nonneg _)) fun i _ => hcm i -/-- If `f` satisfies a boundedness property around `0`, one can deduce a bound on `f m₁ - f m₂` -using the multilinearity. Here, we give a precise but hard to use version. See -`norm_image_sub_le_of_bound` for a less precise but more usable version. The bound reads +/-- If a multilinear map `f` satisfies a boundedness property around `0`, +one can deduce a bound on `f m₁ - f m₂` using the multilinearity. +Here, we give a precise but hard to use version. +See `norm_image_sub_le_of_bound` for a less precise but more usable version. +The bound reads `‖f m - f m'‖ ≤ C * ‖m 1 - m' 1‖ * max ‖m 2‖ ‖m' 2‖ * max ‖m 3‖ ‖m' 3‖ * ... * max ‖m n‖ ‖m' n‖ + ...`, where the other terms in the sum are the same products where `1` is replaced by any `i`. -/ -theorem norm_image_sub_le_of_bound' [DecidableEq ι] {C : ℝ} (hC : 0 ≤ C) +theorem norm_image_sub_le_of_bound' [DecidableEq ι] (f : MultilinearMap 𝕜 E G) {C : ℝ} (hC : 0 ≤ C) (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) (m₁ m₂ : ∀ i, E i) : ‖f m₁ - f m₂‖ ≤ C * ∑ i, ∏ j, if j = i then ‖m₁ i - m₂ i‖ else max ‖m₁ j‖ ‖m₂ j‖ := by have A : @@ -243,8 +244,8 @@ theorem norm_image_sub_le_of_bound' [DecidableEq ι] {C : ℝ} (hC : 0 ≤ C) using the multilinearity. Here, we give a usable but not very precise version. See `norm_image_sub_le_of_bound'` for a more precise but less usable version. The bound is `‖f m - f m'‖ ≤ C * card ι * ‖m - m'‖ * (max ‖m‖ ‖m'‖) ^ (card ι - 1)`. -/ -theorem norm_image_sub_le_of_bound {C : ℝ} (hC : 0 ≤ C) (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) - (m₁ m₂ : ∀ i, E i) : +theorem norm_image_sub_le_of_bound (f : MultilinearMap 𝕜 E G) + {C : ℝ} (hC : 0 ≤ C) (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) (m₁ m₂ : ∀ i, E i) : ‖f m₁ - f m₂‖ ≤ C * Fintype.card ι * max ‖m₁‖ ‖m₂‖ ^ (Fintype.card ι - 1) * ‖m₁ - m₂‖ := by classical have A : @@ -277,7 +278,8 @@ theorem norm_image_sub_le_of_bound {C : ℝ} (hC : 0 ≤ C) (H : ∀ m, ‖f m /-- If a multilinear map satisfies an inequality `‖f m‖ ≤ C * ∏ i, ‖m i‖`, then it is continuous. -/ -theorem continuous_of_bound (C : ℝ) (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) : Continuous f := by +theorem continuous_of_bound (f : MultilinearMap 𝕜 E G) (C : ℝ) (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) : + Continuous f := by let D := max C 1 have D_pos : 0 ≤ D := le_trans zero_le_one (le_max_right _ _) replace H (m) : ‖f m‖ ≤ D * ∏ i, ‖m i‖ := @@ -296,17 +298,19 @@ theorem continuous_of_bound (C : ℝ) (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m /-- Constructing a continuous multilinear map from a multilinear map satisfying a boundedness condition. -/ -def mkContinuous (C : ℝ) (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) : ContinuousMultilinearMap 𝕜 E G := +def mkContinuous (f : MultilinearMap 𝕜 E G) (C : ℝ) (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) : + ContinuousMultilinearMap 𝕜 E G := { f with cont := f.continuous_of_bound C H } @[simp] -theorem coe_mkContinuous (C : ℝ) (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) : ⇑(f.mkContinuous C H) = f := +theorem coe_mkContinuous (f : MultilinearMap 𝕜 E G) (C : ℝ) (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) : + ⇑(f.mkContinuous C H) = f := rfl /-- Given a multilinear map in `n` variables, if one restricts it to `k` variables putting `z` on the other coordinates, then the resulting restricted function satisfies an inequality `‖f.restr v‖ ≤ C * ‖z‖^(n-k) * Π ‖v i‖` if the original function satisfies `‖f v‖ ≤ C * Π ‖v i‖`. -/ -theorem restr_norm_le {k n : ℕ} (f : (MultilinearMap 𝕜 (fun _ : Fin n => G) G' : _)) +theorem restr_norm_le {k n : ℕ} (f : MultilinearMap 𝕜 (fun _ : Fin n => G) G') (s : Finset (Fin n)) (hk : #s = k) (z : G) {C : ℝ} (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) (v : Fin k → G) : ‖f.restr s hk z v‖ ≤ C * ‖z‖ ^ (n - k) * ∏ i, ‖v i‖ := by rw [mul_right_comm, mul_assoc] @@ -326,18 +330,18 @@ smallest number such that `‖f m‖ ≤ ‖f‖ * ∏ i, ‖m i‖` for all `m` defines a normed space structure on `ContinuousMultilinearMap 𝕜 E G`. -/ - namespace ContinuousMultilinearMap -variable [Fintype ι] (c : 𝕜) (f g : ContinuousMultilinearMap 𝕜 E G) (m : ∀ i, E i) +variable [Fintype ι] -theorem bound : ∃ C : ℝ, 0 < C ∧ ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖ := +theorem bound (f : ContinuousMultilinearMap 𝕜 E G) : + ∃ C : ℝ, 0 < C ∧ ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖ := f.toMultilinearMap.exists_bound_of_continuous f.2 open Real /-- The operator norm of a continuous multilinear map is the inf of all its bounds. -/ -def opNorm := +def opNorm (f : ContinuousMultilinearMap 𝕜 E G) : ℝ := sInf { c | 0 ≤ (c : ℝ) ∧ ∀ m, ‖f m‖ ≤ c * ∏ i, ‖m i‖ } instance hasOpNorm : Norm (ContinuousMultilinearMap 𝕜 E G) := @@ -348,7 +352,8 @@ search. -/ instance hasOpNorm' : Norm (ContinuousMultilinearMap 𝕜 (fun _ : ι => G) G') := ContinuousMultilinearMap.hasOpNorm -theorem norm_def : ‖f‖ = sInf { c | 0 ≤ (c : ℝ) ∧ ∀ m, ‖f m‖ ≤ c * ∏ i, ‖m i‖ } := +theorem norm_def (f : ContinuousMultilinearMap 𝕜 E G) : + ‖f‖ = sInf { c | 0 ≤ (c : ℝ) ∧ ∀ m, ‖f m‖ ≤ c * ∏ i, ‖m i‖ } := rfl -- So that invocations of `le_csInf` make sense: we show that the set of @@ -362,7 +367,8 @@ theorem bounds_bddBelow {f : ContinuousMultilinearMap 𝕜 E G} : BddBelow { c | 0 ≤ c ∧ ∀ m, ‖f m‖ ≤ c * ∏ i, ‖m i‖ } := ⟨0, fun _ ⟨hn, _⟩ => hn⟩ -theorem isLeast_opNorm : IsLeast {c : ℝ | 0 ≤ c ∧ ∀ m, ‖f m‖ ≤ c * ∏ i, ‖m i‖} ‖f‖ := by +theorem isLeast_opNorm (f : ContinuousMultilinearMap 𝕜 E G) : + IsLeast {c : ℝ | 0 ≤ c ∧ ∀ m, ‖f m‖ ≤ c * ∏ i, ‖m i‖} ‖f‖ := by refine IsClosed.isLeast_csInf ?_ bounds_nonempty bounds_bddBelow simp only [Set.setOf_and, Set.setOf_forall] exact isClosed_Ici.inter (isClosed_iInter fun m ↦ @@ -370,35 +376,37 @@ 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‖ := +theorem opNorm_nonneg (f : ContinuousMultilinearMap 𝕜 E G) : 0 ≤ ‖f‖ := Real.sInf_nonneg fun _ ⟨hx, _⟩ => hx @[deprecated (since := "2024-02-02")] alias op_norm_nonneg := opNorm_nonneg /-- The fundamental property of the operator norm of a continuous multilinear map: `‖f m‖` is bounded by `‖f‖` times the product of the `‖m i‖`. -/ -theorem le_opNorm : ‖f m‖ ≤ ‖f‖ * ∏ i, ‖m i‖ := f.isLeast_opNorm.1.2 m +theorem le_opNorm (f : ContinuousMultilinearMap 𝕜 E G) (m : ∀ i, E i) : + ‖f m‖ ≤ ‖f‖ * ∏ i, ‖m i‖ := + f.isLeast_opNorm.1.2 m @[deprecated (since := "2024-02-02")] alias le_op_norm := le_opNorm -variable {f m} - -theorem le_mul_prod_of_le_opNorm_of_le {C : ℝ} {b : ι → ℝ} (hC : ‖f‖ ≤ C) (hm : ∀ i, ‖m i‖ ≤ b i) : +theorem le_mul_prod_of_opNorm_le_of_le {f : ContinuousMultilinearMap 𝕜 E G} + {m : ∀ i, E i} {C : ℝ} {b : ι → ℝ} (hC : ‖f‖ ≤ C) (hm : ∀ i, ‖m i‖ ≤ b i) : ‖f m‖ ≤ C * ∏ i, b i := - (f.le_opNorm m).trans <| mul_le_mul hC (prod_le_prod (fun _ _ ↦ norm_nonneg _) fun _ _ ↦ hm _) - (by positivity) ((opNorm_nonneg _).trans hC) + (f.le_opNorm m).trans <| by gcongr; exacts [f.opNorm_nonneg.trans hC, hm _] @[deprecated (since := "2024-02-02")] -alias le_mul_prod_of_le_op_norm_of_le := le_mul_prod_of_le_opNorm_of_le +alias le_mul_prod_of_le_op_norm_of_le := le_mul_prod_of_opNorm_le_of_le -variable (f) +@[deprecated (since := "2024-11-27")] +alias le_mul_prod_of_le_opNorm_of_le := le_mul_prod_of_opNorm_le_of_le -theorem le_opNorm_mul_prod_of_le {b : ι → ℝ} (hm : ∀ i, ‖m i‖ ≤ b i) : ‖f m‖ ≤ ‖f‖ * ∏ i, b i := - le_mul_prod_of_le_opNorm_of_le le_rfl hm +theorem le_opNorm_mul_prod_of_le (f : ContinuousMultilinearMap 𝕜 E G) + {m : ∀ i, E i} {b : ι → ℝ} (hm : ∀ i, ‖m i‖ ≤ b i) : ‖f m‖ ≤ ‖f‖ * ∏ i, b i := + le_mul_prod_of_opNorm_le_of_le le_rfl hm @[deprecated (since := "2024-02-02")] alias le_op_norm_mul_prod_of_le := le_opNorm_mul_prod_of_le -theorem le_opNorm_mul_pow_card_of_le {b : ℝ} (hm : ‖m‖ ≤ b) : +theorem le_opNorm_mul_pow_card_of_le (f : ContinuousMultilinearMap 𝕜 E G) {m b} (hm : ‖m‖ ≤ b) : ‖f m‖ ≤ ‖f‖ * b ^ Fintype.card ι := by simpa only [prod_const] using f.le_opNorm_mul_prod_of_le fun i => (norm_le_pi_norm m i).trans hm @@ -412,47 +420,48 @@ theorem le_opNorm_mul_pow_of_le {n : ℕ} {Ei : Fin n → Type*} [∀ i, Seminor @[deprecated (since := "2024-02-02")] alias le_op_norm_mul_pow_of_le := le_opNorm_mul_pow_of_le -variable {f} (m) - -theorem le_of_opNorm_le {C : ℝ} (h : ‖f‖ ≤ C) : ‖f m‖ ≤ C * ∏ i, ‖m i‖ := - le_mul_prod_of_le_opNorm_of_le h fun _ ↦ le_rfl +theorem le_of_opNorm_le {f : ContinuousMultilinearMap 𝕜 E G} {C : ℝ} (h : ‖f‖ ≤ C) (m : ∀ i, E i) : + ‖f m‖ ≤ C * ∏ i, ‖m i‖ := + le_mul_prod_of_opNorm_le_of_le h fun _ ↦ le_rfl @[deprecated (since := "2024-02-02")] alias le_of_op_norm_le := le_of_opNorm_le -variable (f) - -theorem ratio_le_opNorm : (‖f m‖ / ∏ i, ‖m i‖) ≤ ‖f‖ := +theorem ratio_le_opNorm (f : ContinuousMultilinearMap 𝕜 E G) (m : ∀ i, E i) : + (‖f m‖ / ∏ i, ‖m i‖) ≤ ‖f‖ := div_le_of_le_mul₀ (by positivity) (opNorm_nonneg _) (f.le_opNorm m) @[deprecated (since := "2024-02-02")] alias ratio_le_op_norm := ratio_le_opNorm /-- The image of the unit ball under a continuous multilinear map is bounded. -/ -theorem unit_le_opNorm (h : ‖m‖ ≤ 1) : ‖f m‖ ≤ ‖f‖ := +theorem unit_le_opNorm (f : ContinuousMultilinearMap 𝕜 E G) {m : ∀ i, E i} (h : ‖m‖ ≤ 1) : + ‖f m‖ ≤ ‖f‖ := (le_opNorm_mul_pow_card_of_le f h).trans <| by simp @[deprecated (since := "2024-02-02")] alias unit_le_op_norm := unit_le_opNorm /-- If one controls the norm of every `f x`, then one controls the norm of `f`. -/ -theorem opNorm_le_bound {M : ℝ} (hMp : 0 ≤ M) (hM : ∀ m, ‖f m‖ ≤ M * ∏ i, ‖m i‖) : ‖f‖ ≤ M := +theorem opNorm_le_bound {f : ContinuousMultilinearMap 𝕜 E G} + {M : ℝ} (hMp : 0 ≤ M) (hM : ∀ m, ‖f m‖ ≤ M * ∏ i, ‖m i‖) : ‖f‖ ≤ M := csInf_le bounds_bddBelow ⟨hMp, hM⟩ @[deprecated (since := "2024-02-02")] alias op_norm_le_bound := opNorm_le_bound -theorem opNorm_le_iff {C : ℝ} (hC : 0 ≤ C) : ‖f‖ ≤ C ↔ ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖ := - ⟨fun h _ ↦ le_of_opNorm_le _ h, opNorm_le_bound _ hC⟩ +theorem opNorm_le_iff {f : ContinuousMultilinearMap 𝕜 E G} {C : ℝ} (hC : 0 ≤ C) : + ‖f‖ ≤ C ↔ ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖ := + ⟨fun h _ ↦ le_of_opNorm_le h _, opNorm_le_bound hC⟩ @[deprecated (since := "2024-02-02")] alias op_norm_le_iff := opNorm_le_iff /-- The operator norm satisfies the triangle inequality. -/ -theorem opNorm_add_le : ‖f + g‖ ≤ ‖f‖ + ‖g‖ := - opNorm_le_bound _ (add_nonneg (opNorm_nonneg _) (opNorm_nonneg _)) fun x => by +theorem opNorm_add_le (f g : ContinuousMultilinearMap 𝕜 E G) : ‖f + g‖ ≤ ‖f‖ + ‖g‖ := + opNorm_le_bound (add_nonneg (opNorm_nonneg f) (opNorm_nonneg g)) fun x => by rw [add_mul] exact norm_add_le_of_le (le_opNorm _ _) (le_opNorm _ _) @[deprecated (since := "2024-02-02")] alias op_norm_add_le := opNorm_add_le theorem opNorm_zero : ‖(0 : ContinuousMultilinearMap 𝕜 E G)‖ = 0 := - (opNorm_nonneg _).antisymm' <| opNorm_le_bound 0 le_rfl fun m => by simp + (opNorm_nonneg _).antisymm' <| opNorm_le_bound le_rfl fun m => by simp @[deprecated (since := "2024-02-02")] alias op_norm_zero := opNorm_zero @@ -460,21 +469,13 @@ section variable {𝕜' : Type*} [NormedField 𝕜'] [NormedSpace 𝕜' G] [SMulCommClass 𝕜 𝕜' G] -theorem opNorm_smul_le (c : 𝕜') : ‖c • f‖ ≤ ‖c‖ * ‖f‖ := +theorem opNorm_smul_le (c : 𝕜') (f : ContinuousMultilinearMap 𝕜 E G) : ‖c • f‖ ≤ ‖c‖ * ‖f‖ := (c • f).opNorm_le_bound (mul_nonneg (norm_nonneg _) (opNorm_nonneg _)) fun m ↦ by rw [smul_apply, norm_smul, mul_assoc] exact mul_le_mul_of_nonneg_left (le_opNorm _ _) (norm_nonneg _) @[deprecated (since := "2024-02-02")] alias op_norm_smul_le := opNorm_smul_le -theorem opNorm_neg : ‖-f‖ = ‖f‖ := by - rw [norm_def] - apply congr_arg - ext - simp - -@[deprecated (since := "2024-02-02")] alias op_norm_neg := opNorm_neg - variable (𝕜 E G) in /-- Operator seminorm on the space of continuous multilinear maps, as `Seminorm`. @@ -483,7 +484,7 @@ to define a `SeminormedAddCommGroup` structure on `ContinuousMultilinearMap 𝕜 but we have to override the projection `UniformSpace` so that it is definitionally equal to the one coming from the topologies on `E` and `G`. -/ protected def seminorm : Seminorm 𝕜 (ContinuousMultilinearMap 𝕜 E G) := - .ofSMulLE norm opNorm_zero opNorm_add_le fun c f ↦ opNorm_smul_le f c + .ofSMulLE norm opNorm_zero opNorm_add_le fun c f ↦ f.opNorm_smul_le c private lemma uniformity_eq_seminorm : 𝓤 (ContinuousMultilinearMap 𝕜 E G) = ⨅ r > 0, 𝓟 {f | ‖f.1 - f.2‖ < r} := by @@ -497,7 +498,7 @@ private lemma uniformity_eq_seminorm : suffices ∀ f : ContinuousMultilinearMap 𝕜 E G, (∀ x, ‖x‖ ≤ ‖c‖ → ‖f x‖ ≤ 1) → ‖f‖ ≤ 1 by simpa [NormedSpace.isVonNBounded_closedBall, closedBall_mem_nhds, Set.subset_def, Set.MapsTo] intro f hf - refine opNorm_le_bound _ (by positivity) <| + refine opNorm_le_bound (by positivity) <| f.1.bound_of_shell_of_continuous f.2 (fun _ ↦ hc₀) (fun _ ↦ hc) fun x hcx hx ↦ ?_ calc ‖f x‖ ≤ 1 := hf _ <| (pi_norm_le_iff_of_nonneg (norm_nonneg c)).2 fun i ↦ (hx i).le @@ -540,26 +541,35 @@ search. -/ instance normedSpace' : NormedSpace 𝕜' (ContinuousMultilinearMap 𝕜 (fun _ : ι => G') G) := ContinuousMultilinearMap.normedSpace +@[deprecated (since := "2024-11-24")] +theorem opNorm_neg (f : ContinuousMultilinearMap 𝕜 E G) : ‖-f‖ = ‖f‖ := norm_neg f + +@[deprecated (since := "2024-02-02")] alias op_norm_neg := norm_neg + /-- The fundamental property of the operator norm of a continuous multilinear map: `‖f m‖` is bounded by `‖f‖` times the product of the `‖m i‖`, `nnnorm` version. -/ -theorem le_opNNNorm : ‖f m‖₊ ≤ ‖f‖₊ * ∏ i, ‖m i‖₊ := +theorem le_opNNNorm (f : ContinuousMultilinearMap 𝕜 E G) (m : ∀ i, E i) : + ‖f m‖₊ ≤ ‖f‖₊ * ∏ i, ‖m i‖₊ := NNReal.coe_le_coe.1 <| by push_cast exact f.le_opNorm m @[deprecated (since := "2024-02-02")] alias le_op_nnnorm := le_opNNNorm -theorem le_of_opNNNorm_le {C : ℝ≥0} (h : ‖f‖₊ ≤ C) : ‖f m‖₊ ≤ C * ∏ i, ‖m i‖₊ := +theorem le_of_opNNNorm_le (f : ContinuousMultilinearMap 𝕜 E G) + {C : ℝ≥0} (h : ‖f‖₊ ≤ C) (m : ∀ i, E i) : ‖f m‖₊ ≤ C * ∏ i, ‖m i‖₊ := (f.le_opNNNorm m).trans <| mul_le_mul' h le_rfl @[deprecated (since := "2024-02-02")] alias le_of_op_nnnorm_le := le_of_opNNNorm_le -theorem opNNNorm_le_iff {C : ℝ≥0} : ‖f‖₊ ≤ C ↔ ∀ m, ‖f m‖₊ ≤ C * ∏ i, ‖m i‖₊ := by - simp only [← NNReal.coe_le_coe]; simp [opNorm_le_iff _ C.coe_nonneg, NNReal.coe_prod] +theorem opNNNorm_le_iff {f : ContinuousMultilinearMap 𝕜 E G} {C : ℝ≥0} : + ‖f‖₊ ≤ C ↔ ∀ m, ‖f m‖₊ ≤ C * ∏ i, ‖m i‖₊ := by + simp only [← NNReal.coe_le_coe]; simp [opNorm_le_iff C.coe_nonneg, NNReal.coe_prod] @[deprecated (since := "2024-02-02")] alias op_nnnorm_le_iff := opNNNorm_le_iff -theorem isLeast_opNNNorm : IsLeast {C : ℝ≥0 | ∀ m, ‖f m‖₊ ≤ C * ∏ i, ‖m i‖₊} ‖f‖₊ := by +theorem isLeast_opNNNorm (f : ContinuousMultilinearMap 𝕜 E G) : + IsLeast {C : ℝ≥0 | ∀ m, ‖f m‖₊ ≤ C * ∏ i, ‖m i‖₊} ‖f‖₊ := by simpa only [← opNNNorm_le_iff] using isLeast_Ici @[deprecated (since := "2024-02-02")] alias isLeast_op_nnnorm := isLeast_opNNNorm @@ -629,7 +639,7 @@ variable {G} (E) @[simp] theorem norm_constOfIsEmpty [IsEmpty ι] (x : G) : ‖constOfIsEmpty 𝕜 E x‖ = ‖x‖ := by apply le_antisymm - · refine opNorm_le_bound _ (norm_nonneg _) fun x => ?_ + · refine opNorm_le_bound (norm_nonneg _) fun x => ?_ rw [Fintype.prod_empty, mul_one, constOfIsEmpty_apply] · simpa using (constOfIsEmpty 𝕜 E x).le_opNorm 0 @@ -644,6 +654,7 @@ section variable (𝕜 E E' G G') /-- `ContinuousMultilinearMap.prod` as a `LinearIsometryEquiv`. -/ +@[simps] def prodL : ContinuousMultilinearMap 𝕜 E G × ContinuousMultilinearMap 𝕜 E G' ≃ₗᵢ[𝕜] ContinuousMultilinearMap 𝕜 E (G × G') where @@ -658,11 +669,11 @@ def prodL : norm_map' f := opNorm_prod f.1 f.2 /-- `ContinuousMultilinearMap.pi` as a `LinearIsometryEquiv`. -/ +@[simps! apply symm_apply] def piₗᵢ {ι' : Type v'} [Fintype ι'] {E' : ι' → Type wE'} [∀ i', NormedAddCommGroup (E' i')] [∀ i', NormedSpace 𝕜 (E' i')] : - @LinearIsometryEquiv 𝕜 𝕜 _ _ (RingHom.id 𝕜) _ _ _ (∀ i', ContinuousMultilinearMap 𝕜 E (E' i')) - (ContinuousMultilinearMap 𝕜 E (∀ i, E' i)) _ _ (@Pi.module ι' _ 𝕜 _ _ fun _ => inferInstance) - _ where + (∀ i', ContinuousMultilinearMap 𝕜 E (E' i')) ≃ₗᵢ[𝕜] + ContinuousMultilinearMap 𝕜 E (∀ i, E' i) where toLinearEquiv := piLinearEquiv norm_map' := opNorm_pi @@ -677,7 +688,9 @@ variable [NormedSpace 𝕜' G] [IsScalarTower 𝕜' 𝕜 G] variable [∀ i, NormedSpace 𝕜' (E i)] [∀ i, IsScalarTower 𝕜' 𝕜 (E i)] @[simp] -theorem norm_restrictScalars : ‖f.restrictScalars 𝕜'‖ = ‖f‖ := rfl +theorem norm_restrictScalars (f : ContinuousMultilinearMap 𝕜 E G) : + ‖f.restrictScalars 𝕜'‖ = ‖f‖ := + rfl variable (𝕜') @@ -695,14 +708,14 @@ For a less precise but more usable version, see `norm_image_sub_le`. The bound r `‖f m - f m'‖ ≤ ‖f‖ * ‖m 1 - m' 1‖ * max ‖m 2‖ ‖m' 2‖ * max ‖m 3‖ ‖m' 3‖ * ... * max ‖m n‖ ‖m' n‖ + ...`, where the other terms in the sum are the same products where `1` is replaced by any `i`. -/ -theorem norm_image_sub_le' [DecidableEq ι] (m₁ m₂ : ∀ i, E i) : +theorem norm_image_sub_le' [DecidableEq ι] (f : ContinuousMultilinearMap 𝕜 E G) (m₁ m₂ : ∀ i, E i) : ‖f m₁ - f m₂‖ ≤ ‖f‖ * ∑ i, ∏ j, if j = i then ‖m₁ i - m₂ i‖ else max ‖m₁ j‖ ‖m₂ j‖ := f.toMultilinearMap.norm_image_sub_le_of_bound' (norm_nonneg _) f.le_opNorm _ _ /-- The difference `f m₁ - f m₂` is controlled in terms of `‖f‖` and `‖m₁ - m₂‖`, less precise version. For a more precise but less usable version, see `norm_image_sub_le'`. The bound is `‖f m - f m'‖ ≤ ‖f‖ * card ι * ‖m - m'‖ * (max ‖m‖ ‖m'‖) ^ (card ι - 1)`. -/ -theorem norm_image_sub_le (m₁ m₂ : ∀ i, E i) : +theorem norm_image_sub_le (f : ContinuousMultilinearMap 𝕜 E G) (m₁ m₂ : ∀ i, E i) : ‖f m₁ - f m₂‖ ≤ ‖f‖ * Fintype.card ι * max ‖m₁‖ ‖m₂‖ ^ (Fintype.card ι - 1) * ‖m₁ - m₂‖ := f.toMultilinearMap.norm_image_sub_le_of_bound (norm_nonneg _) f.le_opNorm _ _ @@ -715,14 +728,14 @@ variable [Fintype ι] nonnegative. -/ theorem MultilinearMap.mkContinuous_norm_le (f : MultilinearMap 𝕜 E G) {C : ℝ} (hC : 0 ≤ C) (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) : ‖f.mkContinuous C H‖ ≤ C := - ContinuousMultilinearMap.opNorm_le_bound _ hC fun m => H m + ContinuousMultilinearMap.opNorm_le_bound hC fun m => H m /-- If a continuous multilinear map is constructed from a multilinear map via the constructor `mkContinuous`, then its norm is bounded by the bound given to the constructor if it is nonnegative. -/ theorem MultilinearMap.mkContinuous_norm_le' (f : MultilinearMap 𝕜 E G) {C : ℝ} (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) : ‖f.mkContinuous C H‖ ≤ max C 0 := - ContinuousMultilinearMap.opNorm_le_bound _ (le_max_right _ _) fun m ↦ (H m).trans <| + ContinuousMultilinearMap.opNorm_le_bound (le_max_right _ _) fun m ↦ (H m).trans <| mul_le_mul_of_nonneg_right (le_max_left _ _) <| by positivity namespace ContinuousMultilinearMap @@ -748,7 +761,7 @@ variable {A : Type*} [NormedCommRing A] [NormedAlgebra 𝕜 A] @[simp] theorem norm_mkPiAlgebra_le [Nonempty ι] : ‖ContinuousMultilinearMap.mkPiAlgebra 𝕜 ι A‖ ≤ 1 := by - refine opNorm_le_bound _ zero_le_one fun m => ?_ + refine opNorm_le_bound zero_le_one fun m => ?_ simp only [ContinuousMultilinearMap.mkPiAlgebra_apply, one_mul] exact norm_prod_le' _ univ_nonempty _ @@ -775,7 +788,7 @@ section variable {n : ℕ} {A : Type*} [NormedRing A] [NormedAlgebra 𝕜 A] theorem norm_mkPiAlgebraFin_succ_le : ‖ContinuousMultilinearMap.mkPiAlgebraFin 𝕜 n.succ A‖ ≤ 1 := by - refine opNorm_le_bound _ zero_le_one fun m => ?_ + refine opNorm_le_bound zero_le_one fun m => ?_ 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 ?_ @@ -790,7 +803,7 @@ theorem norm_mkPiAlgebraFin_le_of_pos (hn : 0 < n) : theorem norm_mkPiAlgebraFin_zero : ‖ContinuousMultilinearMap.mkPiAlgebraFin 𝕜 0 A‖ = ‖(1 : A)‖ := by refine le_antisymm ?_ ?_ - · refine opNorm_le_bound _ (norm_nonneg (1 : A)) ?_ + · refine opNorm_le_bound (norm_nonneg (1 : A)) ?_ simp · convert ratio_le_opNorm (ContinuousMultilinearMap.mkPiAlgebraFin 𝕜 0 A) fun _ => (1 : A) simp @@ -818,7 +831,7 @@ end theorem nnnorm_smulRight (f : ContinuousMultilinearMap 𝕜 E 𝕜) (z : G) : ‖f.smulRight z‖₊ = ‖f‖₊ * ‖z‖₊ := by refine le_antisymm ?_ ?_ - · refine (opNNNorm_le_iff _ |>.2 fun m => (nnnorm_smul_le _ _).trans ?_) + · refine opNNNorm_le_iff.2 fun m => (nnnorm_smul_le _ _).trans ?_ rw [mul_right_comm] gcongr exact le_opNNNorm _ _ @@ -893,14 +906,13 @@ namespace ContinuousLinearMap theorem norm_compContinuousMultilinearMap_le (g : G →L[𝕜] G') (f : ContinuousMultilinearMap 𝕜 E G) : ‖g.compContinuousMultilinearMap f‖ ≤ ‖g‖ * ‖f‖ := - ContinuousMultilinearMap.opNorm_le_bound _ (mul_nonneg (norm_nonneg _) (norm_nonneg _)) fun m => + ContinuousMultilinearMap.opNorm_le_bound (by positivity) fun m ↦ calc ‖g (f m)‖ ≤ ‖g‖ * (‖f‖ * ∏ i, ‖m i‖) := g.le_opNorm_of_le <| f.le_opNorm _ _ = _ := (mul_assoc _ _ _).symm variable (𝕜 E G G') - /-- `ContinuousLinearMap.compContinuousMultilinearMap` as a bundled continuous bilinear map. -/ def compContinuousMultilinearMapL : (G →L[𝕜] G') →L[𝕜] ContinuousMultilinearMap 𝕜 E G →L[𝕜] ContinuousMultilinearMap 𝕜 E G' := @@ -959,7 +971,7 @@ def flipMultilinear (f : G →L[𝕜] ContinuousMultilinearMap 𝕜 E G') : simp only [ContinuousMultilinearMap.smul_apply, map_smul, RingHom.id_apply] } (‖f‖ * ∏ i, ‖m i‖) fun x => by rw [mul_right_comm] - exact (f x).le_of_opNorm_le _ (f.le_opNorm x) + exact (f x).le_of_opNorm_le (f.le_opNorm x) _ map_update_add' := fun m i x y => by ext1 simp only [add_apply, ContinuousMultilinearMap.map_update_add, LinearMap.coe_mk, @@ -1065,7 +1077,7 @@ namespace ContinuousMultilinearMap theorem norm_compContinuousLinearMap_le (g : ContinuousMultilinearMap 𝕜 E₁ G) (f : ∀ i, E i →L[𝕜] E₁ i) : ‖g.compContinuousLinearMap f‖ ≤ ‖g‖ * ∏ i, ‖f i‖ := - opNorm_le_bound _ (by positivity) fun m => + opNorm_le_bound (by positivity) fun m => calc ‖g fun i => f i (m i)‖ ≤ ‖g‖ * ∏ i, ‖f i (m i)‖ := g.le_opNorm _ _ ≤ ‖g‖ * ∏ i, ‖f i‖ * ‖m i‖ := @@ -1077,7 +1089,7 @@ theorem norm_compContinuousLinearMap_le (g : ContinuousMultilinearMap 𝕜 E₁ theorem norm_compContinuous_linearIsometry_le (g : ContinuousMultilinearMap 𝕜 E₁ G) (f : ∀ i, E i →ₗᵢ[𝕜] E₁ i) : ‖g.compContinuousLinearMap fun i => (f i).toContinuousLinearMap‖ ≤ ‖g‖ := by - refine opNorm_le_bound _ (norm_nonneg _) fun m => ?_ + refine opNorm_le_bound (norm_nonneg _) fun m => ?_ apply (g.le_opNorm _).trans _ simp only [ContinuousLinearMap.coe_coe, LinearIsometry.coe_toContinuousLinearMap, LinearIsometry.norm_map, le_rfl] @@ -1187,12 +1199,11 @@ continuous-multilinear in `f₁, ..., fₙ`. -/ noncomputable def compContinuousLinearMapContinuousMultilinear : ContinuousMultilinearMap 𝕜 (fun i ↦ E i →L[𝕜] E₁ i) ((ContinuousMultilinearMap 𝕜 E₁ G) →L[𝕜] ContinuousMultilinearMap 𝕜 E G) := - @MultilinearMap.mkContinuous 𝕜 ι (fun i ↦ E i →L[𝕜] E₁ i) - ((ContinuousMultilinearMap 𝕜 E₁ G) →L[𝕜] ContinuousMultilinearMap 𝕜 E G) _ - (fun _ ↦ ContinuousLinearMap.toSeminormedAddCommGroup) - (fun _ ↦ ContinuousLinearMap.toNormedSpace) _ _ - (compContinuousLinearMapMultilinear 𝕜 E E₁ G) _ 1 - fun f ↦ by simpa using norm_compContinuousLinearMapL_le G f + MultilinearMap.mkContinuous (𝕜 := 𝕜) (E := fun i ↦ E i →L[𝕜] E₁ i) + (G := (ContinuousMultilinearMap 𝕜 E₁ G) →L[𝕜] ContinuousMultilinearMap 𝕜 E G) + (compContinuousLinearMapMultilinear 𝕜 E E₁ G) 1 fun f ↦ by + rw [one_mul] + apply norm_compContinuousLinearMapL_le variable {𝕜 E E₁} @@ -1327,18 +1338,16 @@ variable {𝕜 : Type u} {ι : Type v} {E : ι → Type wE} {G : Type wG} {G' : [NontriviallyNormedField 𝕜] [∀ i, SeminormedAddCommGroup (E i)] [∀ i, NormedSpace 𝕜 (E i)] [NormedAddCommGroup G] [NormedSpace 𝕜 G] [SeminormedAddCommGroup G'] [NormedSpace 𝕜 G'] -variable (f : ContinuousMultilinearMap 𝕜 E G) - /-- A continuous linear map is zero iff its norm vanishes. -/ -theorem opNorm_zero_iff : ‖f‖ = 0 ↔ f = 0 := by - simp [← (opNorm_nonneg f).le_iff_eq, opNorm_le_iff f le_rfl, ContinuousMultilinearMap.ext_iff] +theorem opNorm_zero_iff {f : ContinuousMultilinearMap 𝕜 E G} : ‖f‖ = 0 ↔ f = 0 := by + simp [← (opNorm_nonneg f).le_iff_eq, opNorm_le_iff le_rfl, ContinuousMultilinearMap.ext_iff] @[deprecated (since := "2024-02-02")] alias op_norm_zero_iff := opNorm_zero_iff /-- Continuous multilinear maps themselves form a normed group with respect to the operator norm. -/ instance normedAddCommGroup : NormedAddCommGroup (ContinuousMultilinearMap 𝕜 E G) := - NormedAddCommGroup.ofSeparation (fun f ↦ (opNorm_zero_iff f).mp) + NormedAddCommGroup.ofSeparation fun _ ↦ opNorm_zero_iff.mp /-- An alias of `ContinuousMultilinearMap.normedAddCommGroup` with non-dependent types to help typeclass search. -/ @@ -1349,11 +1358,12 @@ instance normedAddCommGroup' : variable (𝕜 G) theorem norm_ofSubsingleton_id [Subsingleton ι] [Nontrivial G] (i : ι) : - ‖ofSubsingleton 𝕜 G G i (.id _ _)‖ = 1 := by simp + ‖ofSubsingleton 𝕜 G G i (.id _ _)‖ = 1 := by + simp theorem nnnorm_ofSubsingleton_id [Subsingleton ι] [Nontrivial G] (i : ι) : ‖ofSubsingleton 𝕜 G G i (.id _ _)‖₊ = 1 := - NNReal.eq <| norm_ofSubsingleton_id _ _ _ + NNReal.eq <| norm_ofSubsingleton_id .. end ContinuousMultilinearMap @@ -1370,12 +1380,11 @@ variable {𝕜 : Type u} {ι : Type v} {E : ι → Type wE} {G : Type wG} [Finty namespace MultilinearMap -variable (f : MultilinearMap 𝕜 E G) - /-- If a multilinear map in finitely many variables on normed spaces satisfies the inequality `‖f m‖ ≤ C * ∏ i, ‖m i‖` on a shell `ε i / ‖c i‖ < ‖m i‖ < ε i` for some positive numbers `ε i` and elements `c i : 𝕜`, `1 < ‖c i‖`, then it satisfies this inequality for all `m`. -/ -theorem bound_of_shell {ε : ι → ℝ} {C : ℝ} (hε : ∀ i, 0 < ε i) {c : ι → 𝕜} (hc : ∀ i, 1 < ‖c i‖) +theorem bound_of_shell (f : MultilinearMap 𝕜 E G) {ε : ι → ℝ} {C : ℝ} {c : ι → 𝕜} + (hε : ∀ i, 0 < ε i) (hc : ∀ i, 1 < ‖c i‖) (hf : ∀ m : ∀ i, E i, (∀ i, ε i / ‖c i‖ ≤ ‖m i‖) → (∀ i, ‖m i‖ < ε i) → ‖f m‖ ≤ C * ∏ i, ‖m i‖) (m : ∀ i, E i) : ‖f m‖ ≤ C * ∏ i, ‖m i‖ := bound_of_shell_of_norm_map_coord_zero f diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean index 31e1077156f18..fd1a31ac0b55a 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean @@ -404,7 +404,7 @@ theorem ContinuousMultilinearMap.fin0_apply_norm (f : G[×0]→L[𝕜] G') {x : obtain rfl : x = 0 := Subsingleton.elim _ _ refine le_antisymm (by simpa using f.le_opNorm 0) ?_ have : ‖ContinuousMultilinearMap.uncurry0 𝕜 G f.curry0‖ ≤ ‖f.curry0‖ := - ContinuousMultilinearMap.opNorm_le_bound _ (norm_nonneg _) fun m => by + ContinuousMultilinearMap.opNorm_le_bound (norm_nonneg _) fun m => by simp [-ContinuousMultilinearMap.apply_zero_uncurry0] simpa [-Matrix.zero_empty] using this @@ -510,7 +510,7 @@ def uncurrySum (f : ContinuousMultilinearMap 𝕜 (fun _ : ι => G) MultilinearMap.mkContinuous (toMultilinearMapLinear.compMultilinearMap f.toMultilinearMap).uncurrySum ‖f‖ fun m => by simpa [Fintype.prod_sum_type, mul_assoc] using - (f (m ∘ Sum.inl)).le_of_opNorm_le (m ∘ Sum.inr) (f.le_opNorm _) + (f (m ∘ Sum.inl)).le_of_opNorm_le (f.le_opNorm _) (m ∘ Sum.inr) @[simp] theorem uncurrySum_apply (f : ContinuousMultilinearMap 𝕜 (fun _ : ι => G) diff --git a/Mathlib/Analysis/NormedSpace/PiTensorProduct/InjectiveSeminorm.lean b/Mathlib/Analysis/NormedSpace/PiTensorProduct/InjectiveSeminorm.lean index 68347d4a673bc..7b0fa337cc06a 100644 --- a/Mathlib/Analysis/NormedSpace/PiTensorProduct/InjectiveSeminorm.lean +++ b/Mathlib/Analysis/NormedSpace/PiTensorProduct/InjectiveSeminorm.lean @@ -430,13 +430,13 @@ protected theorem mapL_smul [DecidableEq ι] (i : ι) (c : 𝕜) (u : E i →L[ Pi.smul_apply] theorem mapL_opNorm : ‖mapL f‖ ≤ ∏ i, ‖f i‖ := by - rw [ContinuousLinearMap.opNorm_le_iff (Finset.prod_nonneg (fun _ _ ↦ norm_nonneg _))] + rw [ContinuousLinearMap.opNorm_le_iff (by positivity)] intro x rw [mapL, liftIsometry] simp only [LinearIsometryEquiv.coe_mk, liftEquiv_apply, LinearMap.mkContinuous_apply] refine le_trans (norm_eval_le_injectiveSeminorm _ _) (mul_le_mul_of_nonneg_right ?_ (norm_nonneg x)) - rw [ContinuousMultilinearMap.opNorm_le_iff _ (Finset.prod_nonneg (fun _ _ ↦ norm_nonneg _))] + rw [ContinuousMultilinearMap.opNorm_le_iff (Finset.prod_nonneg (fun _ _ ↦ norm_nonneg _))] intro m simp only [ContinuousMultilinearMap.compContinuousLinearMap_apply] refine le_trans (injectiveSeminorm_tprod_le (fun i ↦ (f i) (m i))) ?_ From a906a97eefaad7e04aecf27c39b3ece1d9af01aa Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Thu, 28 Nov 2024 16:12:30 +0000 Subject: [PATCH 393/829] feat(Probability): identity, copy, discard and swap kernels (#19481) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add some notable deterministic kernels. These kernels play a particular role in the Markov category point of view on probability transition kernels. From the TestingLowerBounds project. Co-authored-by: Lorenzo Luccioli Co-authored-by: Rémy Degenne --- Mathlib/Probability/Kernel/Basic.lean | 67 +++++++++++++++++++++ Mathlib/Probability/Kernel/Composition.lean | 43 +++++++++++++ 2 files changed, 110 insertions(+) diff --git a/Mathlib/Probability/Kernel/Basic.lean b/Mathlib/Probability/Kernel/Basic.lean index b4ac72bc2f810..6317d5f6bd5e0 100644 --- a/Mathlib/Probability/Kernel/Basic.lean +++ b/Mathlib/Probability/Kernel/Basic.lean @@ -15,6 +15,13 @@ kernels. * `ProbabilityTheory.Kernel.deterministic (f : α → β) (hf : Measurable f)`: kernel `a ↦ Measure.dirac (f a)`. +* `ProbabilityTheory.Kernel.id`: the identity kernel, deterministic kernel for + the identity function. +* `ProbabilityTheory.Kernel.copy α`: the deterministic kernel that maps `x : α` to + the Dirac measure at `(x, x) : α × α`. +* `ProbabilityTheory.Kernel.discard α`: the Markov kernel to the type `Unit`. +* `ProbabilityTheory.Kernel.swap α β`: the deterministic kernel that maps `(x, y)` to + the Dirac measure at `(y, x)`. * `ProbabilityTheory.Kernel.const α (μβ : measure β)`: constant kernel `a ↦ μβ`. * `ProbabilityTheory.Kernel.restrict κ (hs : MeasurableSet s)`: kernel for which the image of `a : α` is `(κ a).restrict s`. @@ -93,6 +100,66 @@ alias set_lintegral_deterministic := setLIntegral_deterministic end Deterministic +section Id + +/-- The identity kernel, that maps `x : α` to the Dirac measure at `x`. -/ +protected noncomputable +def id : Kernel α α := Kernel.deterministic id measurable_id + +instance : IsMarkovKernel (Kernel.id : Kernel α α) := by rw [Kernel.id]; infer_instance + +lemma id_apply (a : α) : Kernel.id a = Measure.dirac a := by + rw [Kernel.id, deterministic_apply, id_def] + +end Id + +section Copy + +/-- The deterministic kernel that maps `x : α` to the Dirac measure at `(x, x) : α × α`. -/ +noncomputable +def copy (α : Type*) [MeasurableSpace α] : Kernel α (α × α) := + Kernel.deterministic (fun x ↦ (x, x)) (measurable_id.prod measurable_id) + +instance : IsMarkovKernel (copy α) := by rw [copy]; infer_instance + +lemma copy_apply (a : α) : copy α a = Measure.dirac (a, a) := by simp [copy, deterministic_apply] + +end Copy + +section Discard + +/-- The Markov kernel to the `Unit` type. -/ +noncomputable +def discard (α : Type*) [MeasurableSpace α] : Kernel α Unit := + Kernel.deterministic (fun _ ↦ ()) measurable_const + +instance : IsMarkovKernel (discard α) := by rw [discard]; infer_instance + +@[simp] +lemma discard_apply (a : α) : discard α a = Measure.dirac () := deterministic_apply _ _ + +end Discard + +section Swap + +/-- The deterministic kernel that maps `(x, y)` to the Dirac measure at `(y, x)`. -/ +noncomputable +def swap (α β : Type*) [MeasurableSpace α] [MeasurableSpace β] : Kernel (α × β) (β × α) := + Kernel.deterministic Prod.swap measurable_swap + +instance : IsMarkovKernel (swap α β) := by rw [swap]; infer_instance + +/-- See `swap_apply'` for a fully applied version of this lemma. -/ +lemma swap_apply (ab : α × β) : swap α β ab = Measure.dirac ab.swap := by + rw [swap, deterministic_apply] + +/-- See `swap_apply` for a partially applied version of this lemma. -/ +lemma swap_apply' (ab : α × β) {s : Set (β × α)} (hs : MeasurableSet s) : + swap α β ab s = s.indicator 1 ab.swap := by + rw [swap_apply, Measure.dirac_apply' _ hs] + +end Swap + section Const /-- Constant kernel, which always returns the same measure. -/ diff --git a/Mathlib/Probability/Kernel/Composition.lean b/Mathlib/Probability/Kernel/Composition.lean index 7227b012d511c..e3b9daa87e50a 100644 --- a/Mathlib/Probability/Kernel/Composition.lean +++ b/Mathlib/Probability/Kernel/Composition.lean @@ -1124,6 +1124,32 @@ lemma deterministic_comp_deterministic (hf : Measurable f) (hg : Measurable g) : (deterministic g hg) ∘ₖ (deterministic f hf) = deterministic (g ∘ f) (hg.comp hf) := by ext; simp [comp_deterministic_eq_comap, comap_apply, deterministic_apply] +@[simp] +lemma comp_id (κ : Kernel α β) : κ ∘ₖ Kernel.id = κ := by + rw [Kernel.id, comp_deterministic_eq_comap, comap_id] + +@[simp] +lemma id_comp (κ : Kernel α β) : Kernel.id ∘ₖ κ = κ := by + rw [Kernel.id, deterministic_comp_eq_map, map_id] + +@[simp] +lemma comp_discard (κ : Kernel α β) [IsMarkovKernel κ] : discard β ∘ₖ κ = discard α := by + ext a s hs; simp [comp_apply' _ _ _ hs] + +@[simp] +lemma swap_copy : (swap α α) ∘ₖ (copy α) = copy α := by + ext a s hs + rw [comp_apply, copy_apply, Measure.dirac_bind (Kernel.measurable _), swap_apply' _ hs, + Measure.dirac_apply' _ hs] + congr + +@[simp] +lemma swap_swap : (swap α β) ∘ₖ (swap β α) = Kernel.id := by + simp_rw [swap, Kernel.deterministic_comp_deterministic, Prod.swap_swap_eq, Kernel.id] + +lemma swap_comp_eq_map {κ : Kernel α (β × γ)} : (swap β γ) ∘ₖ κ = κ.map Prod.swap := by + rw [swap, deterministic_comp_eq_map] + lemma const_comp (μ : Measure γ) (κ : Kernel α β) : const β μ ∘ₖ κ = fun a ↦ (κ a) Set.univ • μ := by ext _ _ hs @@ -1241,12 +1267,29 @@ lemma map_prod_swap (κ : Kernel α β) (η : Kernel α γ) [IsSFiniteKernel κ] refine (lintegral_lintegral_swap ?_).symm exact hf.aemeasurable +@[simp] +lemma swap_prod {κ : Kernel α β} [IsSFiniteKernel κ] {η : Kernel α γ} [IsSFiniteKernel η] : + (swap β γ) ∘ₖ (κ ×ₖ η) = (η ×ₖ κ) := by + rw [swap_comp_eq_map, map_prod_swap] + lemma deterministic_prod_deterministic {f : α → β} {g : α → γ} (hf : Measurable f) (hg : Measurable g) : deterministic f hf ×ₖ deterministic g hg = deterministic (fun a ↦ (f a, g a)) (hf.prod_mk hg) := by ext; simp_rw [prod_apply, deterministic_apply, Measure.dirac_prod_dirac] +lemma compProd_prodMkLeft_eq_comp + (κ : Kernel α β) [IsSFiniteKernel κ] (η : Kernel β γ) [IsSFiniteKernel η] : + κ ⊗ₖ (prodMkLeft α η) = (Kernel.id ×ₖ η) ∘ₖ κ := by + ext a s hs + rw [comp_eq_snd_compProd, compProd_apply hs, snd_apply' _ _ hs, compProd_apply] + swap; · exact measurable_snd hs + simp only [prodMkLeft_apply, Set.mem_setOf_eq, Set.setOf_mem_eq, prod_apply' _ _ _ hs, + id_apply, id_eq] + congr with b + rw [lintegral_dirac'] + exact measurable_measure_prod_mk_left hs + end Prod end Kernel end ProbabilityTheory From 2f728e97cdecd5a36fb58e0f0e7134e2c30083be Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 28 Nov 2024 16:12:32 +0000 Subject: [PATCH 394/829] chore: split Prefunctor out of Quiver (#19525) --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Functor/Basic.lean | 1 + Mathlib/Combinatorics/Quiver/Basic.lean | 78 ----------------- Mathlib/Combinatorics/Quiver/Path.lean | 2 +- Mathlib/Combinatorics/Quiver/Prefunctor.lean | 91 ++++++++++++++++++++ Mathlib/Combinatorics/Quiver/Push.lean | 2 +- 6 files changed, 95 insertions(+), 80 deletions(-) create mode 100644 Mathlib/Combinatorics/Quiver/Prefunctor.lean diff --git a/Mathlib.lean b/Mathlib.lean index b270a2e2785a0..91d08b86a38f8 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2197,6 +2197,7 @@ import Mathlib.Combinatorics.Quiver.Cast import Mathlib.Combinatorics.Quiver.ConnectedComponent import Mathlib.Combinatorics.Quiver.Covering import Mathlib.Combinatorics.Quiver.Path +import Mathlib.Combinatorics.Quiver.Prefunctor import Mathlib.Combinatorics.Quiver.Push import Mathlib.Combinatorics.Quiver.ReflQuiver import Mathlib.Combinatorics.Quiver.SingleObj diff --git a/Mathlib/CategoryTheory/Functor/Basic.lean b/Mathlib/CategoryTheory/Functor/Basic.lean index 4be276945c1e4..8219f68054377 100644 --- a/Mathlib/CategoryTheory/Functor/Basic.lean +++ b/Mathlib/CategoryTheory/Functor/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Tim Baumann, Stephen Morgan, Kim Morrison -/ import Mathlib.CategoryTheory.Category.Basic +import Mathlib.Combinatorics.Quiver.Prefunctor /-! # Functors diff --git a/Mathlib/Combinatorics/Quiver/Basic.lean b/Mathlib/Combinatorics/Quiver/Basic.lean index 955d88dd3b3b9..4b5b56f23e783 100644 --- a/Mathlib/Combinatorics/Quiver/Basic.lean +++ b/Mathlib/Combinatorics/Quiver/Basic.lean @@ -47,84 +47,6 @@ in a quiver or category. -/ infixr:10 " ⟶ " => Quiver.Hom -/-- A morphism of quivers. As we will later have categorical functors extend this structure, -we call it a `Prefunctor`. -/ -structure Prefunctor (V : Type u₁) [Quiver.{v₁} V] (W : Type u₂) [Quiver.{v₂} W] where - /-- The action of a (pre)functor on vertices/objects. -/ - obj : V → W - /-- The action of a (pre)functor on edges/arrows/morphisms. -/ - map : ∀ {X Y : V}, (X ⟶ Y) → (obj X ⟶ obj Y) - -namespace Prefunctor - --- Porting note: added during port. --- 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*} [Quiver V] [Quiver W] {obj : V → W} {map} {X : V} : - (Prefunctor.mk obj map).obj X = obj X := rfl - -lemma mk_map {V W : Type*} [Quiver V] [Quiver W] {obj : V → W} {map} {X Y : V} {f : X ⟶ Y} : - (Prefunctor.mk obj map).map f = map f := rfl - -@[ext (iff := false)] -theorem ext {V : Type u} [Quiver.{v₁} V] {W : Type u₂} [Quiver.{v₂} W] {F G : Prefunctor 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 := by - ext X - apply h_obj - congr - funext X Y f - simpa using h_map X Y f - -/-- The identity morphism between quivers. -/ -@[simps] -def id (V : Type*) [Quiver V] : Prefunctor V V where - obj := fun X => X - map f := f - -instance (V : Type*) [Quiver V] : Inhabited (Prefunctor V V) := - ⟨id V⟩ - -/-- Composition of morphisms between quivers. -/ -@[simps] -def comp {U : Type*} [Quiver U] {V : Type*} [Quiver V] {W : Type*} [Quiver W] - (F : Prefunctor U V) (G : Prefunctor V W) : Prefunctor U W where - obj X := G.obj (F.obj X) - map f := G.map (F.map f) - -@[simp] -theorem comp_id {U V : Type*} [Quiver U] [Quiver V] (F : Prefunctor U V) : - F.comp (id _) = F := rfl - -@[simp] -theorem id_comp {U V : Type*} [Quiver U] [Quiver V] (F : Prefunctor U V) : - (id _).comp F = F := rfl - -@[simp] -theorem comp_assoc {U V W Z : Type*} [Quiver U] [Quiver V] [Quiver W] [Quiver Z] - (F : Prefunctor U V) (G : Prefunctor V W) (H : Prefunctor W Z) : - (F.comp G).comp H = F.comp (G.comp H) := - rfl - -/-- Notation for a prefunctor between quivers. -/ -infixl:50 " ⥤q " => Prefunctor - -/-- Notation for composition of prefunctors. -/ -infixl:60 " ⋙q " => Prefunctor.comp - -/-- Notation for the identity prefunctor on a quiver. -/ -notation "𝟭q" => 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 := by - rw [h] - -end Prefunctor - namespace Quiver /-- `Vᵒᵖ` reverses the direction of all arrows of `V`. -/ diff --git a/Mathlib/Combinatorics/Quiver/Path.lean b/Mathlib/Combinatorics/Quiver/Path.lean index 035044f758b02..e8c3e6c2dc5f5 100644 --- a/Mathlib/Combinatorics/Quiver/Path.lean +++ b/Mathlib/Combinatorics/Quiver/Path.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 David Wärn,. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Wärn, Kim Morrison -/ -import Mathlib.Combinatorics.Quiver.Basic +import Mathlib.Combinatorics.Quiver.Prefunctor import Mathlib.Logic.Lemmas import Batteries.Data.List.Basic diff --git a/Mathlib/Combinatorics/Quiver/Prefunctor.lean b/Mathlib/Combinatorics/Quiver/Prefunctor.lean new file mode 100644 index 0000000000000..98020fcc9bf1b --- /dev/null +++ b/Mathlib/Combinatorics/Quiver/Prefunctor.lean @@ -0,0 +1,91 @@ +/- +Copyright (c) 2021 David Wärn. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Wärn, Kim Morrison +-/ +import Mathlib.Combinatorics.Quiver.Basic + +/-! +# Morphisms of quivers +-/ + +universe v₁ v₂ u u₁ u₂ + +/-- A morphism of quivers. As we will later have categorical functors extend this structure, +we call it a `Prefunctor`. -/ +structure Prefunctor (V : Type u₁) [Quiver.{v₁} V] (W : Type u₂) [Quiver.{v₂} W] where + /-- The action of a (pre)functor on vertices/objects. -/ + obj : V → W + /-- The action of a (pre)functor on edges/arrows/morphisms. -/ + map : ∀ {X Y : V}, (X ⟶ Y) → (obj X ⟶ obj Y) + + +namespace Prefunctor + +-- Porting note: added during port. +-- 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*} [Quiver V] [Quiver W] {obj : V → W} {map} {X : V} : + (Prefunctor.mk obj map).obj X = obj X := rfl + +lemma mk_map {V W : Type*} [Quiver V] [Quiver W] {obj : V → W} {map} {X Y : V} {f : X ⟶ Y} : + (Prefunctor.mk obj map).map f = map f := rfl + +@[ext (iff := false)] +theorem ext {V : Type u} [Quiver.{v₁} V] {W : Type u₂} [Quiver.{v₂} W] {F G : Prefunctor 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 := by + ext X + apply h_obj + congr + funext X Y f + simpa using h_map X Y f + +/-- The identity morphism between quivers. -/ +@[simps] +def id (V : Type*) [Quiver V] : Prefunctor V V where + obj := fun X => X + map f := f + +instance (V : Type*) [Quiver V] : Inhabited (Prefunctor V V) := + ⟨id V⟩ + +/-- Composition of morphisms between quivers. -/ +@[simps] +def comp {U : Type*} [Quiver U] {V : Type*} [Quiver V] {W : Type*} [Quiver W] + (F : Prefunctor U V) (G : Prefunctor V W) : Prefunctor U W where + obj X := G.obj (F.obj X) + map f := G.map (F.map f) + +@[simp] +theorem comp_id {U V : Type*} [Quiver U] [Quiver V] (F : Prefunctor U V) : + F.comp (id _) = F := rfl + +@[simp] +theorem id_comp {U V : Type*} [Quiver U] [Quiver V] (F : Prefunctor U V) : + (id _).comp F = F := rfl + +@[simp] +theorem comp_assoc {U V W Z : Type*} [Quiver U] [Quiver V] [Quiver W] [Quiver Z] + (F : Prefunctor U V) (G : Prefunctor V W) (H : Prefunctor W Z) : + (F.comp G).comp H = F.comp (G.comp H) := + rfl + +/-- Notation for a prefunctor between quivers. -/ +infixl:50 " ⥤q " => Prefunctor + +/-- Notation for composition of prefunctors. -/ +infixl:60 " ⋙q " => Prefunctor.comp + +/-- Notation for the identity prefunctor on a quiver. -/ +notation "𝟭q" => 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 := by + rw [h] + +end Prefunctor diff --git a/Mathlib/Combinatorics/Quiver/Push.lean b/Mathlib/Combinatorics/Quiver/Push.lean index e1c0737afb946..90a1c1778262c 100644 --- a/Mathlib/Combinatorics/Quiver/Push.lean +++ b/Mathlib/Combinatorics/Quiver/Push.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Rémi Bottinelli. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémi Bottinelli -/ -import Mathlib.Combinatorics.Quiver.Basic +import Mathlib.Combinatorics.Quiver.Prefunctor /-! From 52e329815d8c2be3b3630cfef7ae62f68d836fc5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Thu, 28 Nov 2024 16:12:33 +0000 Subject: [PATCH 395/829] feat(Probability): lemmas about composition-product of a measure and a kernel (#19562) From the TestingLowerBounds project. Co-authored-by: Lorenzo Luccioli --- .../Probability/Kernel/MeasureCompProd.lean | 39 ++++++++++++++----- 1 file changed, 30 insertions(+), 9 deletions(-) diff --git a/Mathlib/Probability/Kernel/MeasureCompProd.lean b/Mathlib/Probability/Kernel/MeasureCompProd.lean index b2e9841c037ba..9a7d41a2b61ca 100644 --- a/Mathlib/Probability/Kernel/MeasureCompProd.lean +++ b/Mathlib/Probability/Kernel/MeasureCompProd.lean @@ -25,7 +25,7 @@ This operation, denoted by `⊗ₘ`, takes `μ : Measure α` and `κ : Kernel α open scoped ENNReal -open ProbabilityTheory +open ProbabilityTheory Set namespace MeasureTheory.Measure @@ -58,13 +58,18 @@ lemma compProd_apply [SFinite μ] [IsSFiniteKernel κ] {s : Set (α × β)} (hs simp_rw [compProd, Kernel.compProd_apply hs, Kernel.const_apply, Kernel.prodMkLeft_apply'] rfl +@[simp] +lemma compProd_apply_univ [SFinite μ] [IsMarkovKernel κ] : (μ ⊗ₘ κ) univ = μ univ := by + rw [compProd_apply MeasurableSet.univ] + simp + lemma compProd_apply_prod [SFinite μ] [IsSFiniteKernel κ] {s : Set α} {t : Set β} (hs : MeasurableSet s) (ht : MeasurableSet t) : (μ ⊗ₘ κ) (s ×ˢ t) = ∫⁻ a in s, κ a t ∂μ := by rw [compProd_apply (hs.prod ht), ← lintegral_indicator hs] congr with a classical - rw [Set.indicator_apply] + rw [indicator_apply] split_ifs with ha <;> simp [ha] lemma compProd_congr [IsSFiniteKernel κ] [IsSFiniteKernel η] @@ -91,6 +96,14 @@ lemma ae_compProd_iff [SFinite μ] [IsSFiniteKernel κ] {p : α × β → Prop} (∀ᵐ x ∂(μ ⊗ₘ κ), p x) ↔ ∀ᵐ a ∂μ, ∀ᵐ b ∂(κ a), p (a, b) := Kernel.ae_compProd_iff hp +/-- The composition product of a measure and a constant kernel is the product between the two +measures. -/ +@[simp] +lemma compProd_const {ν : Measure β} [SFinite μ] [SFinite ν] : + μ ⊗ₘ (Kernel.const α ν) = μ.prod ν := by + ext s hs + simp_rw [compProd_apply hs, prod_apply hs, Kernel.const_apply] + lemma compProd_add_left (μ ν : Measure α) [SFinite μ] [SFinite ν] (κ : Kernel α β) : (μ + ν) ⊗ₘ κ = μ ⊗ₘ κ + ν ⊗ₘ κ := by by_cases hκ : IsSFiniteKernel κ @@ -104,6 +117,17 @@ lemma compProd_add_right (μ : Measure α) (κ η : Kernel α β) · simp_rw [Measure.compProd, Kernel.prodMkLeft_add, Kernel.compProd_add_right, Kernel.add_apply] · simp [compProd_of_not_sfinite _ _ hμ] +@[simp] +lemma fst_compProd (μ : Measure α) [SFinite μ] (κ : Kernel α β) [IsMarkovKernel κ] : + (μ ⊗ₘ κ).fst = μ := by + ext s + rw [compProd, Measure.fst, ← Kernel.fst_apply, Kernel.fst_compProd, Kernel.const_apply] + +lemma compProd_smul_left (a : ℝ≥0∞) [SFinite μ] [IsSFiniteKernel κ] : + (a • μ) ⊗ₘ κ = a • (μ ⊗ₘ κ) := by + ext s hs + simp only [compProd_apply hs, lintegral_smul_measure, smul_apply, smul_eq_mul] + section Integral lemma lintegral_compProd [SFinite μ] [IsSFiniteKernel κ] @@ -163,11 +187,8 @@ lemma dirac_unit_compProd_const (μ : Measure β) [IsFiniteMeasure μ] : Measure.dirac () ⊗ₘ Kernel.const Unit μ = μ.map (Prod.mk ()) := by rw [dirac_unit_compProd, Kernel.const_apply] -@[simp] lemma snd_dirac_unit_compProd_const (μ : Measure β) [IsFiniteMeasure μ] : - snd (Measure.dirac () ⊗ₘ Kernel.const Unit μ) = μ := by - rw [dirac_unit_compProd_const, snd, map_map measurable_snd measurable_prod_mk_left] - simp + snd (Measure.dirac () ⊗ₘ Kernel.const Unit μ) = μ := by simp instance : SFinite (μ ⊗ₘ κ) := by rw [compProd]; infer_instance @@ -210,19 +231,19 @@ lemma absolutelyContinuous_of_compProd [SFinite μ] [IsSFiniteKernel κ] [h_zero (h : μ ⊗ₘ κ ≪ ν ⊗ₘ η) : μ ≪ ν := by refine Measure.AbsolutelyContinuous.mk (fun s hs hs0 ↦ ?_) - have h1 : (ν ⊗ₘ η) (s ×ˢ Set.univ) = 0 := by + have h1 : (ν ⊗ₘ η) (s ×ˢ univ) = 0 := by by_cases hν : SFinite ν swap; · simp [compProd_of_not_sfinite _ _ hν] by_cases hη : IsSFiniteKernel η swap; · simp [compProd_of_not_isSFiniteKernel _ _ hη] rw [Measure.compProd_apply_prod hs MeasurableSet.univ] exact setLIntegral_measure_zero _ _ hs0 - have h2 : (μ ⊗ₘ κ) (s ×ˢ Set.univ) = 0 := h h1 + have h2 : (μ ⊗ₘ κ) (s ×ˢ univ) = 0 := h h1 rw [Measure.compProd_apply_prod hs MeasurableSet.univ, lintegral_eq_zero_iff] at h2 swap; · exact Kernel.measurable_coe _ MeasurableSet.univ by_contra hμs have : Filter.NeBot (ae (μ.restrict s)) := by simp [hμs] - obtain ⟨a, ha⟩ : ∃ a, κ a Set.univ = 0 := h2.exists + obtain ⟨a, ha⟩ : ∃ a, κ a univ = 0 := h2.exists refine absurd ha ?_ simp only [Measure.measure_univ_eq_zero] exact (h_zero a).out From 10a37a905cacb4ce8699d05ee155f2fff0c3214c Mon Sep 17 00:00:00 2001 From: Kexing Ying Date: Thu, 28 Nov 2024 16:50:54 +0000 Subject: [PATCH 396/829] feat(MeasureTheory/Measure/MutuallySingular): Disjoint iff mutually singular (#17995) Co-authored-by: Kexing Ying --- Mathlib/Analysis/SpecificLimits/Basic.lean | 8 +- .../MeasureTheory/Measure/MeasureSpace.lean | 64 ++++++++++ .../Measure/MeasureSpaceDef.lean | 2 - .../Measure/MutuallySingular.lean | 114 +++++++++++++++++- 4 files changed, 184 insertions(+), 4 deletions(-) diff --git a/Mathlib/Analysis/SpecificLimits/Basic.lean b/Mathlib/Analysis/SpecificLimits/Basic.lean index e9bbbad7a0c52..d86c021d2cfec 100644 --- a/Mathlib/Analysis/SpecificLimits/Basic.lean +++ b/Mathlib/Analysis/SpecificLimits/Basic.lean @@ -273,8 +273,14 @@ protected theorem ENNReal.tendsto_pow_atTop_nhds_top_iff {r : ℝ≥0∞} : simp only [ENNReal.tendsto_pow_atTop_nhds_zero_iff, inv_zero] at obs simpa [← ENNReal.inv_pow] using obs <| ENNReal.inv_lt_one.mpr r_gt_one -/-! ### Geometric series -/ +lemma ENNReal.eq_zero_of_le_mul_pow {x r : ℝ≥0∞} {ε : ℝ≥0} (hr : r < 1) + (h : ∀ n : ℕ, x ≤ ε * r ^ n) : x = 0 := by + rw [← nonpos_iff_eq_zero] + refine ge_of_tendsto' (f := fun (n : ℕ) ↦ ε * r ^ n) (x := atTop) ?_ h + rw [← mul_zero (M₀ := ℝ≥0∞) (a := ε)] + exact Tendsto.const_mul (tendsto_pow_atTop_nhds_zero_of_lt_one hr) (Or.inr coe_ne_top) +/-! ### Geometric series -/ section Geometric diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index 3d62ba5101860..92a5dcb5aa47b 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -1015,6 +1015,70 @@ instance instCompleteLattice {_ : MeasurableSpace α} : CompleteLattice (Measure end sInf +lemma inf_apply {s : Set α} (hs : MeasurableSet s) : + (μ ⊓ ν) s = sInf {m | ∃ t, m = μ (t ∩ s) + ν (tᶜ ∩ s)} := by + -- `(μ ⊓ ν) s` is defined as `⊓ (t : ℕ → Set α) (ht : s ⊆ ⋃ n, t n), ∑' n, μ (t n) ⊓ ν (t n)` + rw [← sInf_pair, Measure.sInf_apply hs, OuterMeasure.sInf_apply + (image_nonempty.2 <| insert_nonempty μ {ν})] + refine le_antisymm (le_sInf fun m ⟨t, ht₁⟩ ↦ ?_) (le_iInf₂ fun t' ht' ↦ ?_) + · subst ht₁ + -- We first show `(μ ⊓ ν) s ≤ μ (t ∩ s) + ν (tᶜ ∩ s)` for any `t : Set α` + -- For this, define the sequence `t' : ℕ → Set α` where `t' 0 = t ∩ s`, `t' 1 = tᶜ ∩ s` and + -- `∅` otherwise. Then, we have by construction + -- `(μ ⊓ ν) s ≤ ∑' n, μ (t' n) ⊓ ν (t' n) ≤ μ (t' 0) + ν (t' 1) = μ (t ∩ s) + ν (tᶜ ∩ s)`. + set t' : ℕ → Set α := fun n ↦ if n = 0 then t ∩ s else if n = 1 then tᶜ ∩ s else ∅ with ht' + refine (iInf₂_le t' fun x hx ↦ ?_).trans ?_ + · by_cases hxt : x ∈ t + · refine mem_iUnion.2 ⟨0, ?_⟩ + simp [hx, hxt] + · refine mem_iUnion.2 ⟨1, ?_⟩ + simp [hx, hxt] + · simp only [iInf_image, coe_toOuterMeasure, iInf_pair] + rw [tsum_eq_add_tsum_ite 0, tsum_eq_add_tsum_ite 1, if_neg zero_ne_one.symm, + (tsum_eq_zero_iff ENNReal.summable).2 _, add_zero] + · exact add_le_add (inf_le_left.trans <| by simp [ht']) (inf_le_right.trans <| by simp [ht']) + · simp only [ite_eq_left_iff] + intro n hn₁ hn₀ + simp only [ht', if_neg hn₀, if_neg hn₁, measure_empty, iInf_pair, le_refl, inf_of_le_left] + · simp only [iInf_image, coe_toOuterMeasure, iInf_pair] + -- Conversely, fixing `t' : ℕ → Set α` such that `s ⊆ ⋃ n, t' n`, we construct `t : Set α` + -- for which `μ (t ∩ s) + ν (tᶜ ∩ s) ≤ ∑' n, μ (t' n) ⊓ ν (t' n)`. + -- Denoting `I := {n | μ (t' n) ≤ ν (t' n)}`, we set `t = ⋃ n ∈ I, t' n`. + -- Clearly `μ (t ∩ s) ≤ ∑' n ∈ I, μ (t' n)` and `ν (tᶜ ∩ s) ≤ ∑' n ∉ I, ν (t' n)`, so + -- `μ (t ∩ s) + ν (tᶜ ∩ s) ≤ ∑' n ∈ I, μ (t' n) + ∑' n ∉ I, ν (t' n)` + -- where the RHS equals `∑' n, μ (t' n) ⊓ ν (t' n)` by the choice of `I`. + set t := ⋃ n ∈ {k : ℕ | μ (t' k) ≤ ν (t' k)}, t' n with ht + suffices hadd : μ (t ∩ s) + ν (tᶜ ∩ s) ≤ ∑' n, μ (t' n) ⊓ ν (t' n) by + exact le_trans (sInf_le ⟨t, rfl⟩) hadd + have hle₁ : μ (t ∩ s) ≤ ∑' (n : {k | μ (t' k) ≤ ν (t' k)}), μ (t' n) := + (measure_mono inter_subset_left).trans <| measure_biUnion_le _ (to_countable _) _ + have hcap : tᶜ ∩ s ⊆ ⋃ n ∈ {k | ν (t' k) < μ (t' k)}, t' n := by + simp_rw [ht, compl_iUnion] + refine fun x ⟨hx₁, hx₂⟩ ↦ mem_iUnion₂.2 ?_ + obtain ⟨i, hi⟩ := mem_iUnion.1 <| ht' hx₂ + refine ⟨i, ?_, hi⟩ + by_contra h + simp only [mem_setOf_eq, not_lt] at h + exact mem_iInter₂.1 hx₁ i h hi + have hle₂ : ν (tᶜ ∩ s) ≤ ∑' (n : {k | ν (t' k) < μ (t' k)}), ν (t' n) := + (measure_mono hcap).trans (measure_biUnion_le ν (to_countable {k | ν (t' k) < μ (t' k)}) _) + refine (add_le_add hle₁ hle₂).trans ?_ + have heq : {k | μ (t' k) ≤ ν (t' k)} ∪ {k | ν (t' k) < μ (t' k)} = univ := by + ext k; simp [le_or_lt] + conv in ∑' (n : ℕ), μ (t' n) ⊓ ν (t' n) => rw [← tsum_univ, ← heq] + rw [tsum_union_disjoint (f := fun n ↦ μ (t' n) ⊓ ν (t' n)) ?_ ENNReal.summable ENNReal.summable] + · refine add_le_add (tsum_congr ?_).le (tsum_congr ?_).le + · rw [Subtype.forall] + intro n hn; simpa + · rw [Subtype.forall] + intro n hn + rw [mem_setOf_eq] at hn + simp [le_of_lt hn] + · rw [Set.disjoint_iff] + rintro k ⟨hk₁, hk₂⟩ + rw [mem_setOf_eq] at hk₁ hk₂ + exact False.elim <| hk₂.not_le hk₁ + @[simp] theorem _root_.MeasureTheory.OuterMeasure.toMeasure_top : (⊤ : OuterMeasure α).toMeasure (by rw [OuterMeasure.top_caratheodory]; exact le_top) = diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean b/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean index 11481cc61993f..a05138dd08fce 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean @@ -296,8 +296,6 @@ theorem subset_toMeasurable (μ : Measure α) (s : Set α) : s ⊆ toMeasurable theorem ae_le_toMeasurable : s ≤ᵐ[μ] toMeasurable μ s := HasSubset.Subset.eventuallyLE (subset_toMeasurable _ _) - --(subset_toMeasurable _ _).EventuallyLE - @[simp] theorem measurableSet_toMeasurable (μ : Measure α) (s : Set α) : MeasurableSet (toMeasurable μ s) := by diff --git a/Mathlib/MeasureTheory/Measure/MutuallySingular.lean b/Mathlib/MeasureTheory/Measure/MutuallySingular.lean index 257f1ecc3765f..e23107b1af78d 100644 --- a/Mathlib/MeasureTheory/Measure/MutuallySingular.lean +++ b/Mathlib/MeasureTheory/Measure/MutuallySingular.lean @@ -24,7 +24,7 @@ measure, mutually singular open Set -open MeasureTheory NNReal ENNReal +open MeasureTheory NNReal ENNReal Filter namespace MeasureTheory @@ -153,6 +153,118 @@ lemma _root_.MeasurableEmbedding.mutuallySingular_map {β : Type*} {_ : Measurab · rw [hf.map_apply, hf.injective.preimage_image, hμν.measure_nullSet] · rw [hf.map_apply, Set.preimage_compl, hf.injective.preimage_image, hμν.measure_compl_nullSet] +lemma exists_null_set_measure_lt_of_disjoint (h : Disjoint μ ν) {ε : ℝ≥0} (hε : 0 < ε) : + ∃ s, μ s = 0 ∧ ν sᶜ ≤ 2 * ε := by + have h₁ : (μ ⊓ ν) univ = 0 := le_bot_iff.1 (h (inf_le_left (b := ν)) inf_le_right) ▸ rfl + simp_rw [Measure.inf_apply MeasurableSet.univ, inter_univ] at h₁ + have h₂ : ∀ n : ℕ, ∃ t, μ t + ν tᶜ < ε * (1 / 2) ^ n := by + intro n + obtain ⟨m, ⟨t, ht₁, rfl⟩, hm₂⟩ : + ∃ x ∈ {m | ∃ t, m = μ t + ν tᶜ}, x < ε * (1 / 2 : ℝ≥0∞) ^ n := by + refine exists_lt_of_csInf_lt ⟨ν univ, ∅, by simp⟩ <| h₁ ▸ ENNReal.mul_pos ?_ (by simp) + norm_cast + exact hε.ne.symm + exact ⟨t, hm₂⟩ + choose t ht₂ using h₂ + refine ⟨⋂ n, t n, ?_, ?_⟩ + · refine eq_zero_of_le_mul_pow (by norm_num) + fun n ↦ ((measure_mono <| iInter_subset_of_subset n fun _ ht ↦ ht).trans + (le_add_right le_rfl)).trans (ht₂ n).le + · rw [compl_iInter, (by simp [ENNReal.tsum_mul_left, mul_comm] : + 2 * (ε : ℝ≥0∞) = ∑' (n : ℕ), ε * (1 / 2 : ℝ≥0∞) ^ n)] + refine (measure_iUnion_le _).trans ?_ + exact tsum_le_tsum (fun n ↦ (le_add_left le_rfl).trans (ht₂ n).le) + ENNReal.summable ENNReal.summable + +lemma mutuallySingular_of_disjoint (h : Disjoint μ ν) : μ ⟂ₘ ν := by + have h' (n : ℕ) : ∃ s, μ s = 0 ∧ ν sᶜ ≤ (1 / 2) ^ n := by + convert exists_null_set_measure_lt_of_disjoint h (ε := (1 / 2) ^ (n + 1)) + <| pow_pos (by simp) (n + 1) + push_cast + rw [pow_succ, ← mul_assoc, mul_comm, ← mul_assoc] + norm_cast + rw [div_mul_cancel₀, one_mul] + · push_cast + simp + · simp + choose s hs₂ hs₃ using h' + refine Measure.MutuallySingular.mk (t := (⋃ n, s n)ᶜ) (measure_iUnion_null hs₂) ?_ ?_ + · rw [compl_iUnion] + refine eq_zero_of_le_mul_pow (ε := 1) (by norm_num : (1 / 2 : ℝ≥0∞) < 1) <| fun n ↦ ?_ + rw [ENNReal.coe_one, one_mul] + exact (measure_mono <| iInter_subset_of_subset n fun _ ht ↦ ht).trans (hs₃ n) + · rw [union_compl_self] + +lemma MutuallySingular.disjoint (h : μ ⟂ₘ ν) : Disjoint μ ν := by + have h_bot_iff (ξ : Measure α) : ξ ≤ ⊥ ↔ ξ = 0 := by + rw [le_bot_iff] + rfl + intro ξ hξμ hξν + rw [h_bot_iff] + ext s hs + simp only [Measure.coe_zero, Pi.zero_apply] + rw [← inter_union_compl s h.nullSet, measure_union, add_eq_zero] + · exact ⟨measure_inter_null_of_null_right _ <| absolutelyContinuous_of_le hξμ h.measure_nullSet, + measure_inter_null_of_null_right _ <| absolutelyContinuous_of_le hξν h.measure_compl_nullSet⟩ + · exact Disjoint.mono inter_subset_right inter_subset_right disjoint_compl_right + · exact hs.inter h.measurableSet_nullSet.compl + +lemma MutuallySingular.disjoint_ae (h : μ ⟂ₘ ν) : Disjoint (ae μ) (ae ν) := by + rw [disjoint_iff_inf_le] + intro s _ + refine ⟨s ∪ h.nullSetᶜ, ?_, s ∪ h.nullSet, ?_, ?_⟩ + · rw [mem_ae_iff, compl_union, compl_compl] + exact measure_inter_null_of_null_right _ h.measure_nullSet + · rw [mem_ae_iff, compl_union] + exact measure_inter_null_of_null_right _ h.measure_compl_nullSet + · rw [union_eq_compl_compl_inter_compl, union_eq_compl_compl_inter_compl, + ← compl_union, compl_compl, inter_union_compl, compl_compl] + +lemma disjoint_of_disjoint_ae (h : Disjoint (ae μ) (ae ν)) : Disjoint μ ν := by + rw [disjoint_iff_inf_le] at h ⊢ + refine Measure.le_intro fun s hs _ ↦ ?_ + rw [Measure.inf_apply hs] + have : (⊥ : Measure α) = 0 := rfl + simp only [this, Measure.coe_zero, Pi.zero_apply, nonpos_iff_eq_zero] + specialize h (mem_bot (s := sᶜ)) + rw [mem_inf_iff] at h + obtain ⟨t₁, ht₁, t₂, ht₂, h_eq'⟩ := h + have h_eq : s = t₁ᶜ ∪ t₂ᶜ := by + rw [union_eq_compl_compl_inter_compl, compl_compl, compl_compl, ← h_eq', compl_compl] + rw [mem_ae_iff] at ht₁ ht₂ + refine le_antisymm ?_ zero_le' + refine sInf_le_of_le (a := 0) (b := 0) ?_ le_rfl + rw [h_eq] + refine ⟨t₁ᶜ ∩ t₂, Eq.symm ?_⟩ + rw [add_eq_zero] + constructor + · refine measure_inter_null_of_null_left _ ?_ + exact measure_inter_null_of_null_left _ ht₁ + · rw [compl_inter, compl_compl, union_eq_compl_compl_inter_compl, + union_eq_compl_compl_inter_compl, ← compl_union, compl_compl, compl_compl, inter_comm, + inter_comm t₁, union_comm, inter_union_compl] + exact ht₂ + +lemma mutuallySingular_tfae : List.TFAE + [ μ ⟂ₘ ν, + Disjoint μ ν, + Disjoint (ae μ) (ae ν) ] := by + tfae_have 1 → 2 + | h => h.disjoint + tfae_have 2 → 1 + | h => mutuallySingular_of_disjoint h + tfae_have 1 → 3 + | h => h.disjoint_ae + tfae_have 3 → 2 + | h => disjoint_of_disjoint_ae h + tfae_finish + +lemma mutuallySingular_iff_disjoint : μ ⟂ₘ ν ↔ Disjoint μ ν := + mutuallySingular_tfae.out 0 1 + +lemma mutuallySingular_iff_disjoint_ae : μ ⟂ₘ ν ↔ Disjoint (ae μ) (ae ν) := + mutuallySingular_tfae.out 0 2 + end Measure end MeasureTheory From adcc9a19d65fbff7b2eaade3d56eb3d264317858 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 28 Nov 2024 16:50:56 +0000 Subject: [PATCH 397/829] feat: `S ^ n = S` where `S` is a submonoid (#19259) From GrowthInGroups --- Mathlib/Algebra/Group/Subgroup/Pointwise.lean | 35 ++++++++++++++++--- .../Algebra/Group/Submonoid/Pointwise.lean | 13 ++++++- 2 files changed, 43 insertions(+), 5 deletions(-) diff --git a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean index e2c994a14d331..da9423cd8aba9 100644 --- a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean +++ b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean @@ -50,16 +50,43 @@ lemma op_smul_coe_set [Group G] [SetLike S G] [SubgroupClass S G] {s : S} {a : G MulOpposite.op a • (s : Set G) = s := by ext; simp [Set.mem_smul_set_iff_inv_smul_mem, mul_mem_cancel_right, ha] -@[to_additive (attr := simp, norm_cast)] -lemma coe_mul_coe [SetLike S G] [DivInvMonoid G] [SubgroupClass S G] (H : S) : - H * H = (H : Set G) := by aesop (add simp mem_mul) - @[to_additive (attr := simp, norm_cast)] lemma coe_div_coe [SetLike S G] [DivisionMonoid G] [SubgroupClass S G] (H : S) : H / H = (H : Set G) := by simp [div_eq_mul_inv] variable [Group G] [AddGroup A] {s : Set G} +namespace Set + +open Subgroup + +@[to_additive (attr := simp)] +lemma mul_subgroupClosure (hs : s.Nonempty) : s * closure s = closure s := by + rw [← smul_eq_mul, ← Set.iUnion_smul_set] + have h a (ha : a ∈ s) : a • (closure s : Set G) = closure s := + smul_coe_set <| subset_closure ha + simp +contextual [h, hs] + +open scoped RightActions in +@[to_additive (attr := simp)] +lemma subgroupClosure_mul (hs : s.Nonempty) : closure s * s = closure s := by + rw [← Set.iUnion_op_smul_set] + have h a (ha : a ∈ s) : (closure s : Set G) <• a = closure s := + op_smul_coe_set <| subset_closure ha + simp +contextual [h, hs] + +@[to_additive (attr := simp)] +lemma pow_mul_subgroupClosure (hs : s.Nonempty) : ∀ n, s ^ n * closure s = closure s + | 0 => by simp + | n + 1 => by rw [pow_succ, mul_assoc, mul_subgroupClosure hs, pow_mul_subgroupClosure hs] + +@[to_additive (attr := simp)] +lemma subgroupClosure_mul_pow (hs : s.Nonempty) : ∀ n, closure s * s ^ n = closure s + | 0 => by simp + | n + 1 => by rw [pow_succ', ← mul_assoc, subgroupClosure_mul hs, subgroupClosure_mul_pow hs] + +end Set + namespace Subgroup @[to_additive (attr := simp)] diff --git a/Mathlib/Algebra/Group/Submonoid/Pointwise.lean b/Mathlib/Algebra/Group/Submonoid/Pointwise.lean index 9c5e50a98eb6b..cc8c81e07b75a 100644 --- a/Mathlib/Algebra/Group/Submonoid/Pointwise.lean +++ b/Mathlib/Algebra/Group/Submonoid/Pointwise.lean @@ -54,9 +54,20 @@ way to `Module`). open Set Pointwise -variable {α : Type*} {G : Type*} {M : Type*} {R : Type*} {A : Type*} +variable {α G M R A S : Type*} variable [Monoid M] [AddMonoid A] +@[to_additive (attr := simp, norm_cast)] +lemma coe_mul_coe [SetLike S M] [SubmonoidClass S M] (H : S) : H * H = (H : Set M) := by + aesop (add simp mem_mul) + +set_option linter.unusedVariables false in +@[to_additive (attr := simp)] +lemma coe_set_pow [SetLike S M] [SubmonoidClass S M] : + ∀ {n} (hn : n ≠ 0) (H : S), (H ^ n : Set M) = H + | 1, _, H => by simp + | n + 2, _, H => by rw [pow_succ, coe_set_pow n.succ_ne_zero, coe_mul_coe] + /-! Some lemmas about pointwise multiplication and submonoids. Ideally we put these in `GroupTheory.Submonoid.Basic`, but currently we cannot because that file is imported by this. -/ From b79cbba3f7f201bfd1d86449f69c2c5b06a8e916 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 28 Nov 2024 16:50:57 +0000 Subject: [PATCH 398/829] feat(Order): maximal chains are preserved under order isomorphisms (#19471) From LeanCamCombi --- Mathlib/Order/Chain.lean | 45 ++++++++++++++++++++++++++++++++++++---- 1 file changed, 41 insertions(+), 4 deletions(-) diff --git a/Mathlib/Order/Chain.lean b/Mathlib/Order/Chain.lean index f83b31886b200..b15b044ceabf2 100644 --- a/Mathlib/Order/Chain.lean +++ b/Mathlib/Order/Chain.lean @@ -52,10 +52,12 @@ def SuperChain (s t : Set α) : Prop := def IsMaxChain (s : Set α) : Prop := IsChain r s ∧ ∀ ⦃t⦄, IsChain r t → s ⊆ t → s = t -variable {r} {c c₁ c₂ s t : Set α} {a x y : α} +variable {r} {c c₁ c₂ s t : Set α} {a b x y : α} -theorem isChain_empty : IsChain r ∅ := - Set.pairwise_empty _ +@[simp] lemma IsChain.empty : IsChain r ∅ := pairwise_empty _ +@[simp] lemma IsChain.singleton : IsChain r {a} := pairwise_singleton .. + +@[deprecated (since := "2024-11-25")] alias isChain_empty := IsChain.empty theorem Set.Subsingleton.isChain (hs : s.Subsingleton) : IsChain r s := hs.pairwise _ @@ -78,6 +80,9 @@ protected theorem IsChain.insert (hs : IsChain r s) (ha : ∀ b ∈ s, a ≠ b IsChain r (insert a s) := hs.insert_of_symmetric (fun _ _ => Or.symm) ha +lemma IsChain.pair (h : r a b) : IsChain r {a, b} := + IsChain.singleton.insert fun _ hb _ ↦ .inl <| (eq_of_mem_singleton hb).symm.recOn ‹_› + theorem isChain_univ_iff : IsChain r (univ : Set α) ↔ IsTrichotomous α r := by refine ⟨fun h => ⟨fun a b => ?_⟩, fun h => @isChain_of_trichotomous _ _ h univ⟩ rw [or_left_comm, or_iff_not_imp_left] @@ -136,6 +141,14 @@ theorem IsMaxChain.bot_mem [LE α] [OrderBot α] (h : IsMaxChain (· ≤ ·) s) theorem IsMaxChain.top_mem [LE α] [OrderTop α] (h : IsMaxChain (· ≤ ·) s) : ⊤ ∈ s := (h.2 (h.1.insert fun _ _ _ => Or.inr le_top) <| subset_insert _ _).symm ▸ mem_insert _ _ +lemma IsMaxChain.image {s : β → β → Prop} (e : r ≃r s) {c : Set α} (hc : IsMaxChain r c) : + IsMaxChain s (e '' c) where + left := hc.isChain.image _ _ _ fun _ _ ↦ by exact e.map_rel_iff.2 + right t ht hf := by + rw [← e.coe_fn_toEquiv, ← e.toEquiv.eq_preimage_iff_image_eq, preimage_equiv_eq_image_symm] + exact hc.2 (ht.image _ _ _ fun _ _ ↦ by exact e.symm.map_rel_iff.2) + ((e.toEquiv.subset_symm_image _ _).2 hf) + open Classical in /-- Given a set `s`, if there exists a chain `t` strictly including `s`, then `SuccChain s` is one of these chains. Otherwise it is `s`. -/ @@ -302,11 +315,17 @@ theorem top_mem [OrderTop α] (s : Flag α) : (⊤ : α) ∈ s := theorem bot_mem [OrderBot α] (s : Flag α) : (⊥ : α) ∈ s := s.maxChain.bot_mem +/-- Reinterpret a maximal chain as a flag. -/ +def ofIsMaxChain (c : Set α) (hc : IsMaxChain (· ≤ ·) c) : Flag α := ⟨c, hc.isChain, hc.2⟩ + +@[simp, norm_cast] +lemma coe_ofIsMaxChain (c : Set α) (hc) : ofIsMaxChain c hc = c := rfl + end LE section Preorder -variable [Preorder α] {a b : α} +variable [Preorder α] [Preorder β] {a b : α} {s : Flag α} protected theorem le_or_le (s : Flag α) (ha : a ∈ s) (hb : b ∈ s) : a ≤ b ∨ b ≤ a := s.chain_le.total ha hb @@ -320,6 +339,24 @@ instance [OrderBot α] (s : Flag α) : OrderBot s := instance [BoundedOrder α] (s : Flag α) : BoundedOrder s := Subtype.boundedOrder s.bot_mem s.top_mem +lemma mem_iff_forall_le_or_ge : a ∈ s ↔ ∀ ⦃b⦄, b ∈ s → a ≤ b ∨ b ≤ a := + ⟨fun ha b => s.le_or_le ha, fun hb => + of_not_not fun ha => + Set.ne_insert_of_not_mem _ ‹_› <| + s.maxChain.2 (s.chain_le.insert fun c hc _ => hb hc) <| Set.subset_insert _ _⟩ + +/-- Flags are preserved under order isomorphisms. -/ +def map (e : α ≃o β) : Flag α ≃ Flag β + where + toFun s := ofIsMaxChain _ (s.maxChain.image e) + invFun s := ofIsMaxChain _ (s.maxChain.image e.symm) + left_inv s := ext <| e.symm_image_image s + right_inv s := ext <| e.image_symm_image s + +@[simp, norm_cast] lemma coe_map (e : α ≃o β) (s : Flag α) : ↑(map e s) = e '' s := rfl + +@[simp] lemma symm_map (e : α ≃o β) : (map e).symm = map e.symm := rfl + end Preorder section PartialOrder From 97bd49d581362fd625411bbb66bc92e39be05f3c Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Thu, 28 Nov 2024 17:22:35 +0000 Subject: [PATCH 399/829] feat(NumberTheory/Padics): Mahler's theorem on continuous p-adic functions (#19473) This PR completes the proof of Mahler's theorem, showing that the Mahler functions give a Banach space basis for the continuous functions on Z_p. --- Mathlib/NumberTheory/Padics/MahlerBasis.lean | 153 ++++++++++++++++++- 1 file changed, 145 insertions(+), 8 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/MahlerBasis.lean b/Mathlib/NumberTheory/Padics/MahlerBasis.lean index a6be858db5335..8adaf56467425 100644 --- a/Mathlib/NumberTheory/Padics/MahlerBasis.lean +++ b/Mathlib/NumberTheory/Padics/MahlerBasis.lean @@ -7,8 +7,10 @@ import Mathlib.Algebra.Group.ForwardDiff import Mathlib.Analysis.Normed.Group.Ultra import Mathlib.NumberTheory.Padics.ProperSpace import Mathlib.RingTheory.Binomial +import Mathlib.Topology.Algebra.InfiniteSum.Nonarchimedean import Mathlib.Topology.Algebra.Polynomial -import Mathlib.Topology.ContinuousMap.Compact +import Mathlib.Topology.ContinuousMap.ZeroAtInfty +import Mathlib.Topology.MetricSpace.Ultra.ContinuousMaps /-! # The Mahler basis of continuous functions @@ -16,13 +18,12 @@ import Mathlib.Topology.ContinuousMap.Compact In this file we introduce the Mahler basis function `mahler k`, for `k : ℕ`, which is the unique continuous map `ℤ_[p] → ℚ_[p]` agreeing with `n ↦ n.choose k` for `n ∈ ℕ`. -We also show that for any continuous function `f` on `ℤ_[p]` (valued in a `p`-adic normed space), -the iterated forward differences `Δ^[n] f 0` tend to 0. For this, we follow the argument of -Bojanić [bojanic74]. +Using this, we prove Mahler's theorem, showing that for any any continuous function `f` on `ℤ_[p]` +(valued in a `p`-adic normed space `E`), the Mahler series `x ↦ ∑' k, mahler k x • Δ^[n] f 0` +converges (uniformly) to `f`, and this construction defines a Banach-space isomorphism between +`C(ℤ_[p], E)` and the space of sequences `ℕ → E` tending to 0. -In future PR's, we will show that the Mahler functions give a Banach basis for the space of -continuous maps `ℤ_[p] → ℚ_[p]`, with the basis coefficients of `f` given by the forward differences -`Δ^[n] f 0`. +For this, we follow the argument of Bojanić [bojanic74]. ## References @@ -35,7 +36,9 @@ continuous maps `ℤ_[p] → ℚ_[p]`, with the basis coefficients of `f` given Bojanic -/ -open Finset fwdDiff IsUltrametricDist NNReal Filter Topology +open Finset IsUltrametricDist NNReal Filter + +open scoped fwdDiff ZeroAtInfty Topology variable {p : ℕ} [hp : Fact p.Prime] @@ -237,4 +240,138 @@ lemma fwdDiff_tendsto_zero (f : C(ℤ_[p], E)) : Tendsto (Δ_[1]^[·] f 0) atTop end norm_fwdDiff +section mahler_coeff + +variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℚ_[p] E] + (a : E) (n : ℕ) (x : ℤ_[p]) + +/-- +A single term of a Mahler series, given by the product of the scalar-valued continuous map +`mahler n : ℤ_[p] → ℚ_[p]` with a constant vector in some normed `ℚ_[p]`-vector space. +-/ +noncomputable def mahlerTerm : C(ℤ_[p], E) := (mahler n : C(_, ℚ_[p])) • .const _ a + +lemma mahlerTerm_apply : mahlerTerm a n x = mahler n x • a := by + simp only [mahlerTerm, ContinuousMap.smul_apply', ContinuousMap.const_apply] + +lemma norm_mahlerTerm : ‖(mahlerTerm a n : C(ℤ_[p], E))‖ = ‖a‖ := by + simp only [mahlerTerm, ContinuousMap.norm_smul_const, norm_mahler_eq, one_mul] + +/-- A series of the form considered in Mahler's theorem. -/ +noncomputable def mahlerSeries (a : ℕ → E) : C(ℤ_[p], E) := ∑' n, mahlerTerm (a n) n + +variable [IsUltrametricDist E] [CompleteSpace E] {a : ℕ → E} + +/-- A Mahler series whose coefficients tend to 0 is convergent. -/ +lemma hasSum_mahlerSeries (ha : Tendsto a atTop (𝓝 0)) : + HasSum (fun n ↦ mahlerTerm (a n) n) (mahlerSeries a : C(ℤ_[p], E)) := by + refine (NonarchimedeanAddGroup.summable_of_tendsto_cofinite_zero ?_).hasSum + rw [tendsto_zero_iff_norm_tendsto_zero] at ha ⊢ + simpa only [norm_mahlerTerm, Nat.cofinite_eq_atTop] using ha + +/-- Evaluation of a Mahler series is just the pointwise sum. -/ +lemma mahlerSeries_apply (ha : Tendsto a atTop (𝓝 0)) (x : ℤ_[p]) : + mahlerSeries a x = ∑' n, mahler n x • a n := by + simp only [mahlerSeries, ← ContinuousMap.tsum_apply (hasSum_mahlerSeries ha).summable, + mahlerTerm_apply] + +/-- +The value of a Mahler series at a natural number `n` is given by the finite sum of the first `m` +terms, for any `n ≤ m`. +-/ +lemma mahlerSeries_apply_nat (ha : Tendsto a atTop (𝓝 0)) {m n : ℕ} (hmn : m ≤ n) : + mahlerSeries a (m : ℤ_[p]) = ∑ i in range (n + 1), m.choose i • a i := by + have h_van (i) : m.choose (i + (n + 1)) = 0 := Nat.choose_eq_zero_of_lt (by omega) + have aux : Summable fun i ↦ m.choose (i + (n + 1)) • a (i + (n + 1)) := by + simpa only [h_van, zero_smul] using summable_zero + simp only [mahlerSeries_apply ha, mahler_natCast_eq, Nat.cast_smul_eq_nsmul, add_zero, + ← sum_add_tsum_nat_add' (f := fun i ↦ m.choose i • a i) aux, h_van, zero_smul, tsum_zero] + +/-- +The coefficients of a Mahler series can be recovered from the sum by taking forward differences at +`0`. +-/ +lemma fwdDiff_mahlerSeries (ha : Tendsto a atTop (𝓝 0)) (n) : + Δ_[1]^[n] (mahlerSeries a) (0 : ℤ_[p]) = a n := + calc Δ_[1]^[n] (mahlerSeries a) 0 + -- throw away terms after the n'th + _ = Δ_[1]^[n] (fun k ↦ ∑ j ∈ range (n + 1), k.choose j • (a j)) 0 := by + simp only [fwdDiff_iter_eq_sum_shift, zero_add] + refine Finset.sum_congr rfl fun j hj ↦ ?_ + rw [nsmul_one, nsmul_one, + mahlerSeries_apply_nat ha (Nat.lt_succ.mp <| Finset.mem_range.mp hj), Nat.cast_id] + -- bring `Δ_[1]` inside sum + _ = ∑ j ∈ range (n + 1), Δ_[1]^[n] (fun k ↦ k.choose j • (a j)) 0 := by + simp only [fwdDiff_iter_eq_sum_shift, smul_sum] + rw [sum_comm] + -- bring `Δ_[1]` inside scalar-mult + _ = ∑ j ∈ range (n + 1), (Δ_[1]^[n] (fun k ↦ k.choose j : ℕ → ℤ) 0) • (a j) := by + simp only [fwdDiff_iter_eq_sum_shift, zero_add, sum_smul, smul_assoc, Nat.cast_id, + natCast_zsmul] + -- finish using `fwdDiff_iter_choose_zero` + _ = a n := by + simp only [fwdDiff_iter_choose_zero, ite_smul, one_smul, zero_smul, sum_ite_eq, + Finset.mem_range, lt_add_iff_pos_right, zero_lt_one, ↓reduceIte] + +end mahler_coeff + +section mahler_coeff + +variable {p : ℕ} [hp : Fact p.Prime] {E : Type*} + [NormedAddCommGroup E] [NormedSpace ℚ_[p] E] [IsUltrametricDist E] [CompleteSpace E] + +/-- +**Mahler's theorem**: for any continuous function `f` from `ℤ_[p]` to a `p`-adic Banach space, the +Mahler series with coeffients `n ↦ Δ_[1]^[n] f 0` converges to the original function `f`. +-/ +lemma hasSum_mahler (f : C(ℤ_[p], E)) : HasSum (fun n ↦ mahlerTerm (Δ_[1]^[n] f 0) n) f := by + -- First show `∑' n, mahler_term f n` converges to *something*. + have : HasSum (fun n ↦ mahlerTerm (Δ_[1]^[n] f 0) n) + (mahlerSeries (Δ_[1]^[·] f 0) : C(ℤ_[p], E)) := + hasSum_mahlerSeries (PadicInt.fwdDiff_tendsto_zero f) + -- Now show that the sum of the Mahler terms must equal `f` on a dense set, so it is actually `f`. + convert this using 1 + refine ContinuousMap.coe_injective (PadicInt.denseRange_natCast.equalizer + (map_continuous f) (map_continuous _) (funext fun n ↦ ?_)) + simpa only [Function.comp_apply, mahlerSeries_apply_nat (fwdDiff_tendsto_zero f) le_rfl, + zero_add, sum_apply, Pi.smul_apply, nsmul_one] using (shift_eq_sum_fwdDiff_iter 1 f n 0) + +variable (E) in +/-- +The isometric equivalence from `C(ℤ_[p], E)` to the space of sequences in `E` tending to `0` given +by Mahler's theorem, for `E` a nonarchimedean `ℚ_[p]`-Banach space. +-/ +noncomputable def mahlerEquiv : C(ℤ_[p], E) ≃ₗᵢ[ℚ_[p]] C₀(ℕ, E) where + toFun f := ⟨⟨(Δ_[1]^[·] f 0), continuous_of_discreteTopology⟩, + cocompact_eq_atTop (α := ℕ) ▸ fwdDiff_tendsto_zero f⟩ + invFun a := mahlerSeries a + map_add' f g := by + ext x + simp only [ContinuousMap.coe_add, fwdDiff_iter_add, Pi.add_apply, + ZeroAtInftyContinuousMap.coe_mk, ZeroAtInftyContinuousMap.coe_add] + map_smul' r f := by + ext n + simp only [ContinuousMap.coe_smul, RingHom.id_apply, ZeroAtInftyContinuousMap.coe_mk, + ZeroAtInftyContinuousMap.coe_smul, Pi.smul_apply, fwdDiff_iter_const_smul] + left_inv f := (hasSum_mahler f).tsum_eq + right_inv a := ZeroAtInftyContinuousMap.ext <| + fwdDiff_mahlerSeries (cocompact_eq_atTop (α := ℕ) ▸ zero_at_infty a) + norm_map' f := by + simp only [LinearEquiv.coe_mk, ← ZeroAtInftyContinuousMap.norm_toBCF_eq_norm] + apply le_antisymm + · exact BoundedContinuousFunction.norm_le_of_nonempty.mpr + (fun n ↦ norm_fwdDiff_iter_apply_le 1 f 0 n) + · rw [← (hasSum_mahler f).tsum_eq] + refine (norm_tsum_le _).trans (ciSup_le fun n ↦ ?_) + refine le_trans (le_of_eq ?_) (BoundedContinuousFunction.norm_coe_le_norm _ n) + simp only [ZeroAtInftyContinuousMap.toBCF_toFun, ZeroAtInftyContinuousMap.coe_mk, + norm_mahlerTerm, (hasSum_mahler f).tsum_eq] + +lemma mahlerEquiv_apply (f : C(ℤ_[p], E)) : mahlerEquiv E f = fun n ↦ Δ_[1]^[n] f 0 := rfl + +lemma mahlerEquiv_symm_apply (a : C₀(ℕ, E)) : (mahlerEquiv E).symm a = (mahlerSeries (p := p) a) := + rfl + +end mahler_coeff + end PadicInt From a315e3e56aeeb979335b4f3507ad1e8db2f60edf Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Thu, 28 Nov 2024 17:43:16 +0000 Subject: [PATCH 400/829] feat: expand API for `CFC.posPart` (#19221) --- .../ContinuousFunctionalCalculus/PosPart.lean | 141 ++++++++++++++++-- 1 file changed, 130 insertions(+), 11 deletions(-) diff --git a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/PosPart.lean b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/PosPart.lean index 959183c53b686..6c5d5082a67e1 100644 --- a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/PosPart.lean +++ b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/PosPart.lean @@ -16,6 +16,10 @@ the continuous functional calculus and develops the basic API, including the uni positive and negative parts. -/ +open scoped NNReal + +section NonUnital + variable {A : Type*} [NonUnitalRing A] [Module ℝ A] [SMulCommClass ℝ A A] [IsScalarTower ℝ A A] variable [StarRing A] [TopologicalSpace A] variable [NonUnitalContinuousFunctionalCalculus ℝ (IsSelfAdjoint : A → Prop)] @@ -36,6 +40,12 @@ lemma posPart_def (a : A) : a⁺ = cfcₙ (·⁺ : ℝ → ℝ) a := rfl lemma negPart_def (a : A) : a⁻ = cfcₙ (·⁻ : ℝ → ℝ) a := rfl +@[simp] +lemma posPart_zero : (0 : A)⁺ = 0 := by simp [posPart_def] + +@[simp] +lemma negPart_zero : (0 : A)⁻ = 0 := by simp [negPart_def] + @[simp] lemma posPart_mul_negPart (a : A) : a⁺ * a⁻ = 0 := by rw [posPart_def, negPart_def] @@ -80,6 +90,42 @@ lemma posPart_neg (a : A) : (-a)⁺ = a⁻ := by lemma negPart_neg (a : A) : (-a)⁻ = a⁺ := by rw [← eq_comm, ← sub_eq_zero, ← posPart_neg, neg_neg, sub_self] +section SMul + +variable [StarModule ℝ A] + +@[simp] +lemma posPart_smul {r : ℝ≥0} {a : A} : (r • a)⁺ = r • a⁺ := by + by_cases ha : IsSelfAdjoint a + · simp only [CFC.posPart_def, NNReal.smul_def] + rw [← cfcₙ_comp_smul .., ← cfcₙ_smul ..] + refine cfcₙ_congr fun x hx ↦ ?_ + simp [_root_.posPart_def, mul_max_of_nonneg] + · obtain (rfl | hr) := eq_or_ne r 0 + · simp + · have := (not_iff_not.mpr <| (IsSelfAdjoint.all r).smul_iff hr.isUnit (x := a)) |>.mpr ha + simp [CFC.posPart_def, cfcₙ_apply_of_not_predicate a ha, + cfcₙ_apply_of_not_predicate _ this] + +@[simp] +lemma negPart_smul {r : ℝ≥0} {a : A} : (r • a)⁻ = r • a⁻ := by + simpa using posPart_smul (r := r) (a := -a) + +lemma posPart_smul_of_nonneg {r : ℝ} (hr : 0 ≤ r) {a : A} : (r • a)⁺ = r • a⁺ := + posPart_smul (r := ⟨r, hr⟩) + +lemma posPart_smul_of_nonpos {r : ℝ} (hr : r ≤ 0) {a : A} : (r • a)⁺ = -r • a⁻ := by + nth_rw 1 [← neg_neg r] + rw [neg_smul, ← smul_neg, posPart_smul_of_nonneg (neg_nonneg.mpr hr), posPart_neg] + +lemma negPart_smul_of_nonneg {r : ℝ} (hr : 0 ≤ r) {a : A} : (r • a)⁻ = r • a⁻ := by + conv_lhs => rw [← neg_neg r, neg_smul, negPart_neg, posPart_smul_of_nonpos (by simpa), neg_neg] + +lemma negPart_smul_of_nonpos {r : ℝ} (hr : r ≤ 0) {a : A} : (r • a)⁻ = -r • a⁺ := by + conv_lhs => rw [← neg_neg r, neg_smul, negPart_neg, posPart_smul_of_nonneg (by simpa)] + +end SMul + end Unique variable [PartialOrder A] [StarOrderedRing A] @@ -94,22 +140,46 @@ lemma negPart_nonneg (a : A) : 0 ≤ a⁻ := cfcₙ_nonneg (fun x _ ↦ by positivity) +lemma posPart_eq_of_eq_sub_negPart {a b : A} (hab : a = b - a⁻) (hb : 0 ≤ b := by cfc_tac) : + a⁺ = b := by + have ha := hab.symm ▸ hb.isSelfAdjoint.sub (negPart_nonneg a).isSelfAdjoint + nth_rw 1 [← posPart_sub_negPart a] at hab + simpa using hab + +lemma negPart_eq_of_eq_PosPart_sub {a c : A} (hac : a = a⁺ - c) (hc : 0 ≤ c := by cfc_tac) : + a⁻ = c := by + have ha := hac.symm ▸ (posPart_nonneg a).isSelfAdjoint.sub hc.isSelfAdjoint + nth_rw 1 [← posPart_sub_negPart a] at hac + simpa using hac + +lemma le_posPart {a : A} (ha : IsSelfAdjoint a := by cfc_tac) : a ≤ a⁺ := by + simpa [posPart_sub_negPart a] using sub_le_self a⁺ (negPart_nonneg a) + +lemma neg_negPart_le {a : A} (ha : IsSelfAdjoint a := by cfc_tac) : -a⁻ ≤ a := by + simpa only [posPart_sub_negPart a, ← sub_eq_add_neg] + using le_add_of_nonneg_left (a := -a⁻) (posPart_nonneg a) + variable [NonnegSpectrumClass ℝ A] -lemma eq_posPart_iff (a : A) : a = a⁺ ↔ 0 ≤ a := by +lemma posPart_eq_self (a : A) : a⁺ = a ↔ 0 ≤ a := by refine ⟨fun ha ↦ ha ▸ posPart_nonneg a, fun ha ↦ ?_⟩ - conv_lhs => rw [← cfcₙ_id ℝ a] + conv_rhs => rw [← cfcₙ_id ℝ a] rw [posPart_def] refine cfcₙ_congr (fun x hx ↦ ?_) simpa [_root_.posPart_def] using quasispectrum_nonneg_of_nonneg a ha x hx -lemma negPart_eq_zero_iff (a : A) (ha : IsSelfAdjoint a) : +@[deprecated posPart_eq_self (since := "2024-11-18")] +lemma eq_posPart_iff (a : A) : a = a⁺ ↔ 0 ≤ a := by + rw [eq_comm, posPart_eq_self] + +lemma negPart_eq_zero_iff (a : A) (ha : IsSelfAdjoint a := by cfc_tac) : a⁻ = 0 ↔ 0 ≤ a := by - rw [← eq_posPart_iff] + rw [← posPart_eq_self, eq_comm (b := a)] nth_rw 2 [← posPart_sub_negPart a] simp -lemma eq_negPart_iff (a : A) : a = -a⁻ ↔ a ≤ 0 := by +lemma negPart_eq_neg (a : A) : a⁻ = -a ↔ a ≤ 0 := by + rw [← neg_inj, neg_neg, eq_comm] refine ⟨fun ha ↦ by rw [ha, neg_nonpos]; exact negPart_nonneg a, fun ha ↦ ?_⟩ rw [← neg_nonneg] at ha rw [negPart_def, ← cfcₙ_neg] @@ -121,9 +191,13 @@ lemma eq_negPart_iff (a : A) : a = -a⁻ ↔ a ≤ 0 := by rw [← neg_eq_iff_eq_neg, eq_comm] simpa using quasispectrum_nonneg_of_nonneg _ ha _ hx -lemma posPart_eq_zero_iff (a : A) (ha : IsSelfAdjoint a) : +@[deprecated negPart_eq_neg (since := "2024-11-18")] +lemma eq_negPart_iff (a : A) : a = -a⁻ ↔ a ≤ 0 := by + rw [← neg_inj, neg_neg, eq_comm, negPart_eq_neg] + +lemma posPart_eq_zero_iff (a : A) (ha : IsSelfAdjoint a := by cfc_tac) : a⁺ = 0 ↔ a ≤ 0 := by - rw [← eq_negPart_iff] + rw [← negPart_eq_neg, eq_comm (b := -a)] nth_rw 2 [← posPart_sub_negPart a] simp @@ -140,7 +214,7 @@ open NonUnitalContinuousFunctionalCalculus in precisely `a⁺` and `a⁻`. -/ lemma posPart_negPart_unique {a b c : A} (habc : a = b - c) (hbc : b * c = 0) (hb : 0 ≤ b := by cfc_tac) (hc : 0 ≤ c := by cfc_tac) : - b = a⁺ ∧ c = a⁻ := by + a⁺ = b ∧ a⁻ = c := by /- The key idea is to show that `cfcₙ f a = cfcₙ f b + cfcₙ f (-c)` for all real-valued `f` continuous on the union of the spectra of `a`, `b`, and `-c`. Then apply this to `f = (·⁺)`. The equality holds because both sides constitute star homomorphisms which agree on `f = id` since @@ -152,9 +226,8 @@ lemma posPart_negPart_unique {a b c : A} (habc : a = b - c) (hbc : b * c = 0) /- It suffices to show `b = a⁺` since `a⁺ - a⁻ = a = b - c` -/ rw [and_iff_left_of_imp ?of_b_eq] case of_b_eq => - rw [← posPart_sub_negPart a] at habc rintro rfl - linear_combination (norm := abel1) habc + exact negPart_eq_of_eq_PosPart_sub habc hc /- `s := σₙ ℝ a ∪ σₙ ℝ b ∪ σₙ ℝ (-c)` is compact and each of these sets are subsets of `s`. Moreover, `0 ∈ s`. -/ let s := σₙ ℝ a ∪ σₙ ℝ b ∪ σₙ ℝ (-c) @@ -203,6 +276,7 @@ lemma posPart_negPart_unique {a b c : A} (habc : a = b - c) (hbc : b * c = 0) replace key := congr($key f) simp only [cfcₙHomSuperset_apply, NonUnitalStarAlgHom.coe_mk', NonUnitalAlgHom.coe_mk, ψ, Pi.add_apply, cfcₙHom_eq_cfcₙ_extend (·⁺)] at key + symm calc b = cfcₙ (id : ℝ → ℝ) b + cfcₙ (0 : ℝ → ℝ) (-c) := by simp [cfcₙ_id ℝ b] _ = _ := by @@ -211,7 +285,7 @@ lemma posPart_negPart_unique {a b c : A} (habc : a = b - c) (hbc : b * c = 0) refine cfcₙ_congr fun x hx ↦ Eq.symm ?_ lift x to σₙ ℝ _ using hx simp only [Subtype.val_injective.extend_apply, comp_apply, coe_mk, ContinuousMap.coe_mk, - Subtype.map_coe, id_eq, posPart_eq_self, f, Pi.zero_apply, posPart_eq_zero] + Subtype.map_coe, id_eq, _root_.posPart_eq_self, f, Pi.zero_apply, posPart_eq_zero] · exact quasispectrum_nonneg_of_nonneg b hb x.val x.property · obtain ⟨x, hx⟩ := x simp only [← neg_nonneg] @@ -227,6 +301,51 @@ lemma posPart_negPart_unique {a b c : A} (habc : a = b - c) (hbc : b * c = 0) end CFC +end NonUnital + +section Unital + +namespace CFC + +variable {A : Type*} [Ring A] [Algebra ℝ A] [StarRing A] [TopologicalSpace A] +variable [ContinuousFunctionalCalculus ℝ (IsSelfAdjoint : A → Prop)] +variable [UniqueNonUnitalContinuousFunctionalCalculus ℝ A] + +@[simp] +lemma posPart_one : (1 : A)⁺ = 1 := by + rw [CFC.posPart_def, cfcₙ_eq_cfc] + simp + +@[simp] +lemma negPart_one : (1 : A)⁻ = 0 := by + rw [CFC.negPart_def, cfcₙ_eq_cfc] + simp + +@[simp] +lemma posPart_algebraMap (r : ℝ) : (algebraMap ℝ A r)⁺ = algebraMap ℝ A r⁺ := by + rw [CFC.posPart_def, cfcₙ_eq_cfc] + simp + +@[simp] +lemma negPart_algebraMap (r : ℝ) : (algebraMap ℝ A r)⁻ = algebraMap ℝ A r⁻ := by + rw [CFC.negPart_def, cfcₙ_eq_cfc] + simp + +open NNReal in +@[simp] +lemma posPart_algebraMap_nnreal (r : ℝ≥0) : (algebraMap ℝ≥0 A r)⁺ = algebraMap ℝ≥0 A r := by + rw [CFC.posPart_def, cfcₙ_eq_cfc, IsScalarTower.algebraMap_apply ℝ≥0 ℝ A] + simp + +open NNReal in +@[simp] +lemma posPart_natCast (n : ℕ) : (n : A)⁺ = n := by + rw [← map_natCast (algebraMap ℝ≥0 A), posPart_algebraMap_nnreal] + +end CFC + +end Unital + section SpanNonneg variable {A : Type*} [NonUnitalRing A] [Module ℂ A] [SMulCommClass ℂ A A] [IsScalarTower ℂ A A] From 6262f69dea67e2154c0f732fa3d1f814ff7232e4 Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Thu, 28 Nov 2024 19:04:19 +0000 Subject: [PATCH 401/829] feat(Algebra/Category): `Under.pushout f` is left-exact if `f` is flat. (#19213) --- Mathlib.lean | 2 + .../Algebra/Category/Ring/Under/Limits.lean | 213 ++++++++++++++++++ .../Limits/Preserves/Shapes/Products.lean | 14 ++ Mathlib/RingTheory/TensorProduct/Pi.lean | 58 +++++ 4 files changed, 287 insertions(+) create mode 100644 Mathlib/Algebra/Category/Ring/Under/Limits.lean create mode 100644 Mathlib/RingTheory/TensorProduct/Pi.lean diff --git a/Mathlib.lean b/Mathlib.lean index 91d08b86a38f8..ad4182f451904 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -153,6 +153,7 @@ import Mathlib.Algebra.Category.Ring.Instances import Mathlib.Algebra.Category.Ring.Limits import Mathlib.Algebra.Category.Ring.LinearAlgebra import Mathlib.Algebra.Category.Ring.Under.Basic +import Mathlib.Algebra.Category.Ring.Under.Limits import Mathlib.Algebra.Category.Semigrp.Basic import Mathlib.Algebra.Central.Basic import Mathlib.Algebra.Central.Defs @@ -4522,6 +4523,7 @@ import Mathlib.RingTheory.TensorProduct.Basic import Mathlib.RingTheory.TensorProduct.Finite import Mathlib.RingTheory.TensorProduct.Free import Mathlib.RingTheory.TensorProduct.MvPolynomial +import Mathlib.RingTheory.TensorProduct.Pi import Mathlib.RingTheory.Trace.Basic import Mathlib.RingTheory.Trace.Defs import Mathlib.RingTheory.Trace.Quotient diff --git a/Mathlib/Algebra/Category/Ring/Under/Limits.lean b/Mathlib/Algebra/Category/Ring/Under/Limits.lean new file mode 100644 index 0000000000000..810501a4eca05 --- /dev/null +++ b/Mathlib/Algebra/Category/Ring/Under/Limits.lean @@ -0,0 +1,213 @@ +/- +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.Algebra.Category.Ring.Under.Basic +import Mathlib.CategoryTheory.Limits.Constructions.LimitsOfProductsAndEqualizers +import Mathlib.CategoryTheory.Limits.Over +import Mathlib.RingTheory.TensorProduct.Pi +import Mathlib.RingTheory.Flat.Algebra +import Mathlib.RingTheory.Flat.Equalizer + +/-! +# Limits in `Under R` for a commutative ring `R` + +We show that `Under.pushout f` is left-exact, i.e. preserves finite limits, if `f : R ⟶ S` is +flat. + +-/ + +noncomputable section + +universe v u + +open TensorProduct CategoryTheory Limits + +variable {R S : CommRingCat.{u}} + +namespace CommRingCat.Under + +section Algebra + +variable [Algebra R S] + +section Pi + +variable {ι : Type u} (P : ι → Under R) + +/-- The canonical fan on `P : ι → Under R` given by `∀ i, P i`. -/ +def piFan : Fan P := + Fan.mk (Under.mk <| ofHom <| Pi.ringHom (fun i ↦ (P i).hom)) + (fun i ↦ Under.homMk (Pi.evalRingHom _ i)) + +/-- The canonical fan is limiting. -/ +def piFanIsLimit : IsLimit (piFan P) := + isLimitOfReflects (Under.forget R) <| + (isLimitMapConeFanMkEquiv (Under.forget R) P _).symm <| + CommRingCat.piFanIsLimit (fun i ↦ (P i).right) + +variable (S) in +/-- The fan on `i ↦ S ⊗[R] P i` given by `S ⊗[R] ∀ i, P i` -/ +def tensorProductFan : Fan (fun i ↦ mkUnder S (S ⊗[R] (P i).right)) := + Fan.mk (mkUnder S <| S ⊗[R] ∀ i, (P i).right) + (fun i ↦ AlgHom.toUnder <| + Algebra.TensorProduct.map (AlgHom.id S S) (Pi.evalAlgHom R (fun j ↦ (P j).right) i)) + +variable (S) in +/-- The fan on `i ↦ S ⊗[R] P i` given by `∀ i, S ⊗[R] P i` -/ +def tensorProductFan' : Fan (fun i ↦ mkUnder S (S ⊗[R] (P i).right)) := + Fan.mk (mkUnder S <| ∀ i, S ⊗[R] (P i).right) + (fun i ↦ AlgHom.toUnder <| Pi.evalAlgHom S _ i) + +/-- The two fans on `i ↦ S ⊗[R] P i` agree if `ι` is finite. -/ +def tensorProductFanIso [Fintype ι] [DecidableEq ι] : + tensorProductFan S P ≅ tensorProductFan' S P := + Fan.ext (Algebra.TensorProduct.piRight R S _ _).toUnder <| fun i ↦ by + dsimp only [tensorProductFan, Fan.mk_pt, fan_mk_proj, tensorProductFan'] + apply CommRingCat.mkUnder_ext + intro c + induction c + · simp only [AlgHom.toUnder_right, map_zero, Under.comp_right, comp_apply, + AlgEquiv.toUnder_hom_right_apply, Pi.evalAlgHom_apply, Pi.zero_apply] + · simp only [AlgHom.toUnder_right, Algebra.TensorProduct.map_tmul, AlgHom.coe_id, id_eq, + Pi.evalAlgHom_apply, Under.comp_right, comp_apply, AlgEquiv.toUnder_hom_right_apply, + Algebra.TensorProduct.piRight_tmul] + · simp_all + +open Classical in +/-- The fan on `i ↦ S ⊗[R] P i` given by `S ⊗[R] ∀ i, P i` is limiting if `ι` is finite. -/ +def tensorProductFanIsLimit [Finite ι] : IsLimit (tensorProductFan S P) := + letI : Fintype ι := Fintype.ofFinite ι + (IsLimit.equivIsoLimit (tensorProductFanIso P)).symm (Under.piFanIsLimit _) + +/-- `tensorProd R S` preserves the limit of the canonical fan on `P`. -/ +def piFanTensorProductIsLimit [Finite ι] : IsLimit ((tensorProd R S).mapCone (Under.piFan P)) := + (isLimitMapConeFanMkEquiv (tensorProd R S) P _).symm <| tensorProductFanIsLimit P + +instance (J : Type u) [Finite J] (f : J → Under R) : + PreservesLimit (Discrete.functor f) (tensorProd R S) := + let c : Fan _ := Under.piFan f + have hc : IsLimit c := Under.piFanIsLimit f + preservesLimit_of_preserves_limit_cone hc (piFanTensorProductIsLimit f) + +instance (J : Type) [Finite J] : + PreservesLimitsOfShape (Discrete J) (tensorProd R S) := + let J' : Type u := ULift.{u} J + have : PreservesLimitsOfShape (Discrete J') (tensorProd R S) := + preservesLimitsOfShape_of_discrete (tensorProd R S) + let e : Discrete J' ≌ Discrete J := Discrete.equivalence Equiv.ulift + preservesLimitsOfShape_of_equiv e (R.tensorProd S) + +instance : PreservesFiniteProducts (tensorProd R S) where + preserves J := { } + +end Pi + +section Equalizer + +lemma equalizer_comp {A B : Under R} (f g : A ⟶ B) : + (AlgHom.equalizer (toAlgHom f) (toAlgHom g)).val.toUnder ≫ f = + (AlgHom.equalizer (toAlgHom f) (toAlgHom g)).val.toUnder ≫ g := by + ext (a : AlgHom.equalizer (toAlgHom f) (toAlgHom g)) + exact a.property + +/-- The canonical fork on `f g : A ⟶ B` given by the equalizer. -/ +def equalizerFork {A B : Under R} (f g : A ⟶ B) : + Fork f g := + Fork.ofι ((AlgHom.equalizer (toAlgHom f) (toAlgHom g)).val.toUnder) + (by rw [equalizer_comp]) + +@[simp] +lemma equalizerFork_ι {A B : Under R} (f g : A ⟶ B) : + (Under.equalizerFork f g).ι = (AlgHom.equalizer (toAlgHom f) (toAlgHom g)).val.toUnder := rfl + +/-- Variant of `Under.equalizerFork'` for algebra maps. This is definitionally equal to +`Under.equalizerFork` but this is costly in applications. -/ +def equalizerFork' {A B : Type u} [CommRing A] [CommRing B] [Algebra R A] [Algebra R B] + (f g : A →ₐ[R] B) : + Fork f.toUnder g.toUnder := + Fork.ofι ((AlgHom.equalizer f g).val.toUnder) <| by ext a; exact a.property + +@[simp] +lemma equalizerFork'_ι {A B : Type u} [CommRing A] [CommRing B] [Algebra R A] [Algebra R B] + (f g : A →ₐ[R] B) : + (Under.equalizerFork' f g).ι = (AlgHom.equalizer f g).val.toUnder := rfl + +/-- The canonical fork on `f g : A ⟶ B` is limiting. -/ +def equalizerForkIsLimit {A B : Under R} (f g : A ⟶ B) : + IsLimit (Under.equalizerFork f g) := + isLimitOfReflects (Under.forget R) <| + (isLimitMapConeForkEquiv (Under.forget R) (equalizer_comp f g)).invFun <| + CommRingCat.equalizerForkIsLimit f.right g.right + +/-- Variant of `Under.equalizerForkIsLimit` for algebra maps. -/ +def equalizerFork'IsLimit {A B : Type u} [CommRing A] [CommRing B] [Algebra R A] + [Algebra R B] (f g : A →ₐ[R] B) : + IsLimit (Under.equalizerFork' f g) := + Under.equalizerForkIsLimit f.toUnder g.toUnder + +/-- The fork on `𝟙 ⊗[R] f` and `𝟙 ⊗[R] g` given by `S ⊗[R] eq(f, g)`. -/ +def tensorProdEqualizer {A B : Under R} (f g : A ⟶ B) : + Fork ((tensorProd R S).map f) ((tensorProd R S).map g) := + Fork.ofι + ((tensorProd R S).map ((AlgHom.equalizer (toAlgHom f) (toAlgHom g)).val.toUnder)) <| by + rw [← Functor.map_comp, equalizer_comp, Functor.map_comp] + +@[simp] +lemma tensorProdEqualizer_ι {A B : Under R} (f g : A ⟶ B) : + (tensorProdEqualizer f g).ι = (tensorProd R S).map + ((AlgHom.equalizer (toAlgHom f) (toAlgHom g)).val.toUnder) := + rfl + +/-- If `S` is `R`-flat, `S ⊗[R] eq(f, g)` is isomorphic to `eq(𝟙 ⊗[R] f, 𝟙 ⊗[R] g)`. -/ +def equalizerForkTensorProdIso [Module.Flat R S] {A B : Under R} (f g : A ⟶ B) : + tensorProdEqualizer f g ≅ Under.equalizerFork' + (Algebra.TensorProduct.map (AlgHom.id S S) (toAlgHom f)) + (Algebra.TensorProduct.map (AlgHom.id S S) (toAlgHom g)) := + Fork.ext (AlgHom.tensorEqualizerEquiv S S (toAlgHom f) (toAlgHom g)).toUnder <| by + ext + apply AlgHom.coe_tensorEqualizer + +/-- If `S` is `R`-flat, `tensorProd R S` preserves the equalizer of `f` and `g`. -/ +def tensorProdMapEqualizerForkIsLimit [Module.Flat R S] {A B : Under R} (f g : A ⟶ B) : + IsLimit ((tensorProd R S).mapCone <| Under.equalizerFork f g) := + (isLimitMapConeForkEquiv (tensorProd R S) _).symm <| + (IsLimit.equivIsoLimit (equalizerForkTensorProdIso f g).symm) <| + Under.equalizerFork'IsLimit _ _ + +instance [Module.Flat R S] {A B : Under R} (f g : A ⟶ B) : + PreservesLimit (parallelPair f g) (tensorProd R S) := + let c : Fork f g := Under.equalizerFork f g + let hc : IsLimit c := Under.equalizerForkIsLimit f g + let hc' : IsLimit ((tensorProd R S).mapCone c) := + tensorProdMapEqualizerForkIsLimit f g + preservesLimit_of_preserves_limit_cone hc hc' + +instance [Module.Flat R S] : PreservesLimitsOfShape WalkingParallelPair (tensorProd R S) where + preservesLimit {K} := + preservesLimit_of_iso_diagram _ (diagramIsoParallelPair K).symm + +instance [Module.Flat R S] : PreservesFiniteLimits (tensorProd R S) := + preservesFiniteLimits_of_preservesEqualizers_and_finiteProducts (tensorProd R S) + +end Equalizer + +end Algebra + +variable (f : R ⟶ S) + +/-- `Under.pushout f` preserves finite products. -/ +instance : PreservesFiniteProducts (Under.pushout f) where + preserves _ := + letI : Algebra R S := RingHom.toAlgebra f + preservesLimitsOfShape_of_natIso (tensorProdIsoPushout R S) + +/-- `Under.pushout f` preserves finite limits if `f` is flat. -/ +lemma preservesFiniteLimits_of_flat (hf : RingHom.Flat f) : + PreservesFiniteLimits (Under.pushout f) where + preservesFiniteLimits _ := + letI : Algebra R S := RingHom.toAlgebra f + preservesLimitsOfShape_of_natIso (tensorProdIsoPushout R S) + +end CommRingCat.Under diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Products.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Products.lean index 6f15d54ba6cac..e9b74760f9f31 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Products.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Products.lean @@ -157,4 +157,18 @@ instance : IsIso (sigmaComparison G f) := by end +/-- If `F` preserves the limit of every `Discrete.functor f`, it preserves all limits of shape +`Discrete J`. -/ +lemma preservesLimitsOfShape_of_discrete (F : C ⥤ D) + [∀ (f : J → C), PreservesLimit (Discrete.functor f) F] : + PreservesLimitsOfShape (Discrete J) F where + preservesLimit := preservesLimit_of_iso_diagram F (Discrete.natIsoFunctor).symm + +/-- If `F` preserves the colimit of every `Discrete.functor f`, it preserves all colimits of shape +`Discrete J`. -/ +lemma preservesColimitsOfShape_of_discrete (F : C ⥤ D) + [∀ (f : J → C), PreservesColimit (Discrete.functor f) F] : + PreservesColimitsOfShape (Discrete J) F where + preservesColimit := preservesColimit_of_iso_diagram F (Discrete.natIsoFunctor).symm + end CategoryTheory.Limits diff --git a/Mathlib/RingTheory/TensorProduct/Pi.lean b/Mathlib/RingTheory/TensorProduct/Pi.lean new file mode 100644 index 0000000000000..b6b48e4414f31 --- /dev/null +++ b/Mathlib/RingTheory/TensorProduct/Pi.lean @@ -0,0 +1,58 @@ +/- +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.Algebra.Algebra.Pi +import Mathlib.LinearAlgebra.TensorProduct.Pi +import Mathlib.RingTheory.TensorProduct.Basic + +/-! +# Tensor product and products of algebras + +In this file we examine the behaviour of the tensor product with (finite) products. This +is a direct application of `Mathlib.LinearAlgebra.TensorProduct.Pi` to the algebra case. + +-/ + +open TensorProduct + +section + +variable (R S A : Type*) [CommSemiring R] [CommSemiring S] [Algebra R S] [CommSemiring A] + [Algebra R A] [Algebra S A] [IsScalarTower R S A] +variable {ι : Type*} (B : ι → Type*) [∀ i, CommSemiring (B i)] [∀ i, Algebra R (B i)] + +@[simp] +lemma piRightHom_one : piRightHom R S A B 1 = 1 := rfl + +variable {R S A B} in +@[simp] +lemma piRightHom_mul (x y : A ⊗[R] ∀ i, B i) : + piRightHom R S A B (x * y) = piRightHom R S A B x * piRightHom R S A B y := by + induction x + · simp + · induction y + · simp + · ext j + simp + · simp_all [mul_add] + · simp_all [add_mul] + +/-- The canonical map `A ⊗[R] (∀ i, B i) →ₐ[S] ∀ i, A ⊗[R] B i`. This is an isomorphism +if `ι` is finite (see `Algebra.TensorProduct.piRight`). -/ +noncomputable def piRightHom : A ⊗[R] (∀ i, B i) →ₐ[S] ∀ i, A ⊗[R] B i := + AlgHom.ofLinearMap (_root_.TensorProduct.piRightHom R S A B) (by simp) (by simp) + +variable [Fintype ι] [DecidableEq ι] + +/-- Tensor product of rings commutes with finite products on the right. -/ +noncomputable def Algebra.TensorProduct.piRight : + A ⊗[R] (∀ i, B i) ≃ₐ[S] ∀ i, A ⊗[R] B i := + AlgEquiv.ofLinearEquiv (_root_.TensorProduct.piRight R S A B) (by simp) (by simp) + +@[simp] +lemma Algebra.TensorProduct.piRight_tmul (x : A) (f : ∀ i, B i) : + piRight R S A B (x ⊗ₜ f) = (fun j ↦ x ⊗ₜ f j) := rfl + +end From f60ba8f84d827a2af342e2904775ce632018a06e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 28 Nov 2024 19:53:02 +0000 Subject: [PATCH 402/829] chore: tag all `closure_mono` lemmas `gcongr` (#19561) From GrowthInGroups (LeanCamCombi) --- Mathlib/Algebra/Field/Subfield/Basic.lean | 1 + Mathlib/Algebra/Group/Subgroup/Lattice.lean | 2 +- Mathlib/Algebra/Group/Submonoid/Basic.lean | 2 +- Mathlib/Algebra/Group/Subsemigroup/Basic.lean | 4 ++-- Mathlib/Algebra/Ring/Subring/Basic.lean | 1 + Mathlib/Algebra/Ring/Subsemiring/Basic.lean | 1 + Mathlib/Data/Matroid/Closure.lean | 1 + Mathlib/Deprecated/Subfield.lean | 1 + Mathlib/Deprecated/Subgroup.lean | 2 +- Mathlib/Deprecated/Submonoid.lean | 2 +- Mathlib/Deprecated/Subring.lean | 1 + Mathlib/ModelTheory/Substructures.lean | 1 + Mathlib/RingTheory/NonUnitalSubring/Basic.lean | 1 + Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean | 1 + 14 files changed, 15 insertions(+), 6 deletions(-) diff --git a/Mathlib/Algebra/Field/Subfield/Basic.lean b/Mathlib/Algebra/Field/Subfield/Basic.lean index b4b3a10a096e9..9265de24cc7bb 100644 --- a/Mathlib/Algebra/Field/Subfield/Basic.lean +++ b/Mathlib/Algebra/Field/Subfield/Basic.lean @@ -315,6 +315,7 @@ theorem closure_le {s : Set K} {t : Subfield K} : closure s ≤ t ↔ s ⊆ t := /-- Subfield closure of a set is monotone in its argument: if `s ⊆ t`, then `closure s ≤ closure t`. -/ +@[gcongr] theorem closure_mono ⦃s t : Set K⦄ (h : s ⊆ t) : closure s ≤ closure t := closure_le.2 <| Set.Subset.trans h subset_closure diff --git a/Mathlib/Algebra/Group/Subgroup/Lattice.lean b/Mathlib/Algebra/Group/Subgroup/Lattice.lean index ea029ac82db4b..8d9277bb28c46 100644 --- a/Mathlib/Algebra/Group/Subgroup/Lattice.lean +++ b/Mathlib/Algebra/Group/Subgroup/Lattice.lean @@ -402,7 +402,7 @@ variable {G} /-- Subgroup closure of a set is monotone in its argument: if `h ⊆ k`, then `closure h ≤ closure k`. -/ -@[to_additive +@[to_additive (attr := gcongr) "Additive subgroup closure of a set is monotone in its argument: if `h ⊆ k`, then `closure h ≤ closure k`"] theorem closure_mono ⦃h k : Set G⦄ (h' : h ⊆ k) : closure h ≤ closure k := diff --git a/Mathlib/Algebra/Group/Submonoid/Basic.lean b/Mathlib/Algebra/Group/Submonoid/Basic.lean index 04690c8d30132..db88bd856087e 100644 --- a/Mathlib/Algebra/Group/Submonoid/Basic.lean +++ b/Mathlib/Algebra/Group/Submonoid/Basic.lean @@ -138,7 +138,7 @@ theorem closure_le : closure s ≤ S ↔ s ⊆ S := /-- Submonoid closure of a set is monotone in its argument: if `s ⊆ t`, then `closure s ≤ closure t`. -/ -@[to_additive +@[to_additive (attr := gcongr) "Additive submonoid closure of a set is monotone in its argument: if `s ⊆ t`, then `closure s ≤ closure t`"] theorem closure_mono ⦃s t : Set M⦄ (h : s ⊆ t) : closure s ≤ closure t := diff --git a/Mathlib/Algebra/Group/Subsemigroup/Basic.lean b/Mathlib/Algebra/Group/Subsemigroup/Basic.lean index 72a594ef2bd7a..aafb3813bc835 100644 --- a/Mathlib/Algebra/Group/Subsemigroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subsemigroup/Basic.lean @@ -124,8 +124,8 @@ theorem closure_le : closure s ≤ S ↔ s ⊆ S := /-- subsemigroup closure of a set is monotone in its argument: if `s ⊆ t`, then `closure s ≤ closure t`. -/ -@[to_additive "Additive subsemigroup closure of a set is monotone in its argument: if `s ⊆ t`, - then `closure s ≤ closure t`"] +@[to_additive (attr := gcongr) "Additive subsemigroup closure of a set is monotone in its argument: +if `s ⊆ t`, then `closure s ≤ closure t`"] theorem closure_mono ⦃s t : Set M⦄ (h : s ⊆ t) : closure s ≤ closure t := closure_le.2 <| Subset.trans h subset_closure diff --git a/Mathlib/Algebra/Ring/Subring/Basic.lean b/Mathlib/Algebra/Ring/Subring/Basic.lean index 791ac77d64f46..26820e434059b 100644 --- a/Mathlib/Algebra/Ring/Subring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subring/Basic.lean @@ -460,6 +460,7 @@ theorem closure_le {s : Set R} {t : Subring R} : closure s ≤ t ↔ s ⊆ t := /-- Subring closure of a set is monotone in its argument: if `s ⊆ t`, then `closure s ≤ closure t`. -/ +@[gcongr] theorem closure_mono ⦃s t : Set R⦄ (h : s ⊆ t) : closure s ≤ closure t := closure_le.2 <| Set.Subset.trans h subset_closure diff --git a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean index bcd9bb5829c17..5a0b8a674f389 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean @@ -357,6 +357,7 @@ theorem closure_le {s : Set R} {t : Subsemiring R} : closure s ≤ t ↔ s ⊆ t /-- Subsemiring closure of a set is monotone in its argument: if `s ⊆ t`, then `closure s ≤ closure t`. -/ +@[gcongr] theorem closure_mono ⦃s t : Set R⦄ (h : s ⊆ t) : closure s ≤ closure t := closure_le.2 <| Set.Subset.trans h subset_closure diff --git a/Mathlib/Data/Matroid/Closure.lean b/Mathlib/Data/Matroid/Closure.lean index 4b7a3fddea385..2dbba090c7c1d 100644 --- a/Mathlib/Data/Matroid/Closure.lean +++ b/Mathlib/Data/Matroid/Closure.lean @@ -176,6 +176,7 @@ lemma Flat.closure (hF : M.Flat F) : M.closure F = F := @[simp] lemma closure_univ (M : Matroid α) : M.closure univ = M.E := by rw [← closure_inter_ground, univ_inter, closure_ground] +@[gcongr] lemma closure_subset_closure (M : Matroid α) (h : X ⊆ Y) : M.closure X ⊆ M.closure Y := subset_sInter (fun _ h' ↦ sInter_subset_of_mem ⟨h'.1, subset_trans (inter_subset_inter_left _ h) h'.2⟩) diff --git a/Mathlib/Deprecated/Subfield.lean b/Mathlib/Deprecated/Subfield.lean index f1d508d21d08d..3ce93a3a26acc 100644 --- a/Mathlib/Deprecated/Subfield.lean +++ b/Mathlib/Deprecated/Subfield.lean @@ -125,6 +125,7 @@ theorem closure_subset {T : Set F} (hT : IsSubfield T) (H : S ⊆ T) : closure S theorem closure_subset_iff {s t : Set F} (ht : IsSubfield t) : closure s ⊆ t ↔ s ⊆ t := ⟨Set.Subset.trans subset_closure, closure_subset ht⟩ +@[gcongr] theorem closure_mono {s t : Set F} (H : s ⊆ t) : closure s ⊆ closure t := closure_subset closure.isSubfield <| Set.Subset.trans H subset_closure diff --git a/Mathlib/Deprecated/Subgroup.lean b/Mathlib/Deprecated/Subgroup.lean index 64fd450477d80..745637e6ec963 100644 --- a/Mathlib/Deprecated/Subgroup.lean +++ b/Mathlib/Deprecated/Subgroup.lean @@ -442,7 +442,7 @@ theorem closure_subset {s t : Set G} (ht : IsSubgroup t) (h : s ⊆ t) : closure theorem closure_subset_iff {s t : Set G} (ht : IsSubgroup t) : closure s ⊆ t ↔ s ⊆ t := ⟨fun h _ ha => h (mem_closure ha), fun h _ ha => closure_subset ht h ha⟩ -@[to_additive] +@[to_additive (attr := gcongr)] theorem closure_mono {s t : Set G} (h : s ⊆ t) : closure s ⊆ closure t := closure_subset (closure.isSubgroup _) <| Set.Subset.trans h subset_closure diff --git a/Mathlib/Deprecated/Submonoid.lean b/Mathlib/Deprecated/Submonoid.lean index 639dc73b90592..2e9618d300fb9 100644 --- a/Mathlib/Deprecated/Submonoid.lean +++ b/Mathlib/Deprecated/Submonoid.lean @@ -268,7 +268,7 @@ theorem closure_subset {s t : Set M} (ht : IsSubmonoid t) (h : s ⊆ t) : Closur /-- Given subsets `t` and `s` of a monoid `M`, if `s ⊆ t`, the submonoid of `M` generated by `s` is contained in the submonoid generated by `t`. -/ -@[to_additive +@[to_additive (attr := gcongr) "Given subsets `t` and `s` of an `AddMonoid M`, if `s ⊆ t`, the `AddSubmonoid` of `M` generated by `s` is contained in the `AddSubmonoid` generated by `t`."] theorem closure_mono {s t : Set M} (h : s ⊆ t) : Closure s ⊆ Closure t := diff --git a/Mathlib/Deprecated/Subring.lean b/Mathlib/Deprecated/Subring.lean index a26eec3948fb8..05bea1fc07bb0 100644 --- a/Mathlib/Deprecated/Subring.lean +++ b/Mathlib/Deprecated/Subring.lean @@ -179,6 +179,7 @@ theorem closure_subset_iff {s t : Set R} (ht : IsSubring t) : closure s ⊆ t (AddGroup.closure_subset_iff ht.toIsAddSubgroup).trans ⟨Set.Subset.trans Monoid.subset_closure, Monoid.closure_subset ht.toIsSubmonoid⟩ +@[gcongr] theorem closure_mono {s t : Set R} (H : s ⊆ t) : closure s ⊆ closure t := closure_subset closure.isSubring <| Set.Subset.trans H subset_closure diff --git a/Mathlib/ModelTheory/Substructures.lean b/Mathlib/ModelTheory/Substructures.lean index d5f9680e0bee2..a3b0c3233492f 100644 --- a/Mathlib/ModelTheory/Substructures.lean +++ b/Mathlib/ModelTheory/Substructures.lean @@ -251,6 +251,7 @@ theorem closure_le : closure L s ≤ S ↔ s ⊆ S := /-- Substructure closure of a set is monotone in its argument: if `s ⊆ t`, then `closure L s ≤ closure L t`. -/ +@[gcongr] theorem closure_mono ⦃s t : Set M⦄ (h : s ⊆ t) : closure L s ≤ closure L t := (closure L).monotone h diff --git a/Mathlib/RingTheory/NonUnitalSubring/Basic.lean b/Mathlib/RingTheory/NonUnitalSubring/Basic.lean index b857b42737e1a..e376982889514 100644 --- a/Mathlib/RingTheory/NonUnitalSubring/Basic.lean +++ b/Mathlib/RingTheory/NonUnitalSubring/Basic.lean @@ -391,6 +391,7 @@ theorem closure_le {s : Set R} {t : NonUnitalSubring R} : closure s ≤ t ↔ s /-- `NonUnitalSubring` closure of a set is monotone in its argument: if `s ⊆ t`, then `closure s ≤ closure t`. -/ +@[gcongr] theorem closure_mono ⦃s t : Set R⦄ (h : s ⊆ t) : closure s ≤ closure t := closure_le.2 <| Set.Subset.trans h subset_closure diff --git a/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean b/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean index 71ae33dcc8107..a9859dcb69aec 100644 --- a/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean +++ b/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean @@ -330,6 +330,7 @@ theorem closure_le {s : Set R} {t : NonUnitalSubsemiring R} : closure s ≤ t /-- Subsemiring closure of a set is monotone in its argument: if `s ⊆ t`, then `closure s ≤ closure t`. -/ +@[gcongr] theorem closure_mono ⦃s t : Set R⦄ (h : s ⊆ t) : closure s ≤ closure t := closure_le.2 <| Set.Subset.trans h subset_closure From 4e9fa40d7c480937e09cd6e47a591bd6f3b8be42 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 28 Nov 2024 20:24:35 +0000 Subject: [PATCH 403/829] =?UTF-8?q?feat:=20`s=20=E2=8A=86=20s=20^=20n`=20w?= =?UTF-8?q?hen=20`1=20=E2=88=88=20s`=20(#19538)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit and the product of nontrivial sets is nontrivial. From GrowthInGroups --- .../Algebra/Group/Pointwise/Finset/Basic.lean | 45 ++++++++++++++----- .../Algebra/Group/Pointwise/Set/Basic.lean | 38 ++++++++++++++++ .../Additive/CauchyDavenport.lean | 10 ++--- .../Additive/PluenneckeRuzsa.lean | 2 +- 4 files changed, 78 insertions(+), 17 deletions(-) diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index 876f75a97121b..082727c5e5d61 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -968,6 +968,10 @@ lemma pow_subset_pow_right (hs : 1 ∈ s) (hmn : m ≤ n) : s ^ m ⊆ s ^ n := lemma pow_subset_pow (hst : s ⊆ t) (ht : 1 ∈ t) (hmn : m ≤ n) : s ^ m ⊆ t ^ n := (pow_subset_pow_left hst).trans (pow_subset_pow_right ht hmn) +@[to_additive] +lemma subset_pow (hs : 1 ∈ s) (hn : n ≠ 0) : s ⊆ s ^ n := by + simpa using pow_subset_pow_right hs <| Nat.one_le_iff_ne_zero.2 hn + @[deprecated (since := "2024-11-19")] alias pow_subset_pow_of_one_mem := pow_subset_pow_right @[deprecated (since := "2024-11-19")] @@ -1515,7 +1519,16 @@ end Semigroup section IsLeftCancelMul -variable [Mul α] [IsLeftCancelMul α] [DecidableEq α] (s t : Finset α) (a : α) +variable [Mul α] [IsLeftCancelMul α] [DecidableEq α] {s t : Finset α} {a : α} + +@[to_additive] +lemma Nontrivial.mul_left : t.Nontrivial → s.Nonempty → (s * t).Nontrivial := by + rintro ⟨a, ha, b, hb, hab⟩ ⟨c, hc⟩ + exact ⟨c * a, mul_mem_mul hc ha, c * b, mul_mem_mul hc hb, by simpa⟩ + +@[to_additive] +lemma Nontrivial.mul (hs : s.Nontrivial) (ht : t.Nontrivial) : (s * t).Nontrivial := + ht.mul_left hs.nonempty @[to_additive] theorem pairwiseDisjoint_smul_iff {s : Set α} {t : Finset α} : @@ -1523,11 +1536,11 @@ theorem pairwiseDisjoint_smul_iff {s : Set α} {t : Finset α} : simp_rw [← pairwiseDisjoint_coe, coe_smul_finset, Set.pairwiseDisjoint_smul_iff] @[to_additive (attr := simp)] -theorem card_singleton_mul : ({a} * t).card = t.card := +theorem card_singleton_mul (a : α) (t : Finset α) : ({a} * t).card = t.card := card_image₂_singleton_left _ <| mul_right_injective _ @[to_additive] -theorem singleton_mul_inter : {a} * (s ∩ t) = {a} * s ∩ ({a} * t) := +theorem singleton_mul_inter (a : α) (s t : Finset α) : {a} * (s ∩ t) = {a} * s ∩ ({a} * t) := image₂_singleton_inter _ _ <| mul_right_injective _ @[to_additive] @@ -1547,20 +1560,25 @@ theorem card_le_card_mul_self {s : Finset α} : s.card ≤ (s * s).card := by end IsLeftCancelMul -section +section IsRightCancelMul + +variable [Mul α] [IsRightCancelMul α] [DecidableEq α] {s t : Finset α} {a : α} -variable [Mul α] [IsRightCancelMul α] [DecidableEq α] (s t : Finset α) (a : α) +@[to_additive] +lemma Nontrivial.mul_right : s.Nontrivial → t.Nonempty → (s * t).Nontrivial := by + rintro ⟨a, ha, b, hb, hab⟩ ⟨c, hc⟩ + exact ⟨a * c, mul_mem_mul ha hc, b * c, mul_mem_mul hb hc, by simpa⟩ @[to_additive (attr := simp)] -theorem card_mul_singleton : (s * {a}).card = s.card := +theorem card_mul_singleton (s : Finset α) (a : α) : (s * {a}).card = s.card := card_image₂_singleton_right _ <| mul_left_injective _ @[to_additive] -theorem inter_mul_singleton : s ∩ t * {a} = s * {a} ∩ (t * {a}) := +theorem inter_mul_singleton (s t : Finset α) (a : α) : s ∩ t * {a} = s * {a} ∩ (t * {a}) := image₂_inter_singleton _ _ <| mul_left_injective _ @[to_additive] -theorem card_le_card_mul_right {t : Finset α} (ht : t.Nonempty) : s.card ≤ (s * t).card := +theorem card_le_card_mul_right (ht : t.Nonempty) : s.card ≤ (s * t).card := card_le_card_image₂_right _ ht mul_left_injective /-- @@ -1571,18 +1589,23 @@ See `card_le_card_mul_self` for the version with left-cancellative multiplicatio "The size of `s + s` is at least the size of `s`, version with right-cancellative addition. See `card_le_card_add_self` for the version with left-cancellative addition." ] -theorem card_le_card_mul_self' {s : Finset α} : s.card ≤ (s * s).card := by +theorem card_le_card_mul_self' : s.card ≤ (s * s).card := by cases s.eq_empty_or_nonempty <;> simp [card_le_card_mul_right, *] -end +end IsRightCancelMul section CancelMonoid variable [DecidableEq α] [CancelMonoid α] {s : Finset α} {m n : ℕ} +@[to_additive] +lemma Nontrivial.pow (hs : s.Nontrivial) : ∀ {n}, n ≠ 0 → (s ^ n).Nontrivial + | 1, _ => by simpa + | n + 2, _ => by simpa [pow_succ] using (hs.pow n.succ_ne_zero).mul hs + /-- See `Finset.card_pow_mono` for a version that works for the empty set. -/ @[to_additive "See `Finset.card_nsmul_mono` for a version that works for the empty set."] protected lemma Nonempty.card_pow_mono (hs : s.Nonempty) : Monotone fun n : ℕ ↦ (s ^ n).card := - monotone_nat_of_le_succ fun n ↦ by rw [pow_succ]; exact card_le_card_mul_right _ hs + monotone_nat_of_le_succ fun n ↦ by rw [pow_succ]; exact card_le_card_mul_right hs /-- See `Finset.Nonempty.card_pow_mono` for a version that works for zero powers. -/ @[to_additive "See `Finset.Nonempty.card_nsmul_mono` for a version that works for zero scalars."] diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean index 49a613a4045d4..641e71ad326f6 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean @@ -1097,6 +1097,10 @@ lemma pow_subset_pow_right (hs : 1 ∈ s) (hmn : m ≤ n) : s ^ m ⊆ s ^ n := lemma pow_subset_pow (hst : s ⊆ t) (ht : 1 ∈ t) (hmn : m ≤ n) : s ^ m ⊆ t ^ n := (pow_subset_pow_left hst).trans (pow_subset_pow_right ht hmn) +@[to_additive] +lemma subset_pow (hs : 1 ∈ s) (hn : n ≠ 0) : s ⊆ s ^ n := by + simpa using pow_subset_pow_right hs <| Nat.one_le_iff_ne_zero.2 hn + @[deprecated (since := "2024-11-19")] alias pow_subset_pow_of_one_mem := pow_subset_pow_right @[deprecated (since := "2024-11-19")] @@ -1163,6 +1167,40 @@ protected theorem _root_.IsUnit.set : IsUnit a → IsUnit ({a} : Set α) := end Monoid +section IsLeftCancelMul +variable [Mul α] [IsLeftCancelMul α] {s t : Set α} + +@[to_additive] +lemma Nontrivial.mul_left : t.Nontrivial → s.Nonempty → (s * t).Nontrivial := by + rintro ⟨a, ha, b, hb, hab⟩ ⟨c, hc⟩ + exact ⟨c * a, mul_mem_mul hc ha, c * b, mul_mem_mul hc hb, by simpa⟩ + +@[to_additive] +lemma Nontrivial.mul (hs : s.Nontrivial) (ht : t.Nontrivial) : (s * t).Nontrivial := + ht.mul_left hs.nonempty + +end IsLeftCancelMul + +section IsRightCancelMul +variable [Mul α] [IsRightCancelMul α] {s t : Set α} + +@[to_additive] +lemma Nontrivial.mul_right : s.Nontrivial → t.Nonempty → (s * t).Nontrivial := by + rintro ⟨a, ha, b, hb, hab⟩ ⟨c, hc⟩ + exact ⟨a * c, mul_mem_mul ha hc, b * c, mul_mem_mul hb hc, by simpa⟩ + +end IsRightCancelMul + +section CancelMonoid +variable [CancelMonoid α] {s t : Set α} {a : α} {n : ℕ} + +@[to_additive] +lemma Nontrivial.pow (hs : s.Nontrivial) : ∀ {n}, n ≠ 0 → (s ^ n).Nontrivial + | 1, _ => by simpa + | n + 2, _ => by simpa [pow_succ] using (hs.pow n.succ_ne_zero).mul hs + +end CancelMonoid + /-- `Set α` is a `CommMonoid` under pointwise operations if `α` is. -/ @[to_additive "`Set α` is an `AddCommMonoid` under pointwise operations if `α` is."] protected noncomputable def commMonoid [CommMonoid α] : CommMonoid (Set α) := diff --git a/Mathlib/Combinatorics/Additive/CauchyDavenport.lean b/Mathlib/Combinatorics/Additive/CauchyDavenport.lean index c986855dbedef..ec8c94780e944 100644 --- a/Mathlib/Combinatorics/Additive/CauchyDavenport.lean +++ b/Mathlib/Combinatorics/Additive/CauchyDavenport.lean @@ -100,8 +100,8 @@ private lemma wellFoundedOn_devosMulRel : Set.WellFoundedOn.prod_lex_of_wellFoundedOn_fiber ?_ fun n ↦ wellFounded_lt.onFun.wellFoundedOn exact wellFounded_lt.onFun.wellFoundedOn.mono' fun x hx y _ ↦ tsub_lt_tsub_left_of_le <| - add_le_add ((card_le_card_mul_right _ hx.1.2).trans_eq hx.2) <| - (card_le_card_mul_left _ hx.1.1).trans_eq hx.2 + add_le_add ((card_le_card_mul_right hx.1.2).trans_eq hx.2) <| + (card_le_card_mul_left hx.1.1).trans_eq hx.2 /-- A generalisation of the **Cauchy-Davenport theorem** to arbitrary groups. The size of `s * t` is lower-bounded by `|s| + |t| - 1` unless this quantity is greater than the size of the smallest @@ -151,7 +151,7 @@ lemma cauchy_davenport_minOrder_mul (hs : s.Nonempty) (ht : t.Nonempty) : refine Or.inl ((minOrder_le_natCard (zpowers_ne_bot.2 hg) <| s.finite_toSet.smul_set.subset hS).trans <| WithTop.coe_le_coe.2 <| ((Nat.card_mono s.finite_toSet.smul_set hS).trans_eq <| ?_).trans <| - card_le_card_mul_right _ ht) + card_le_card_mul_right ht) rw [← coe_smul_finset] simp [-coe_smul_finset] -- Else, we can transform `s`, `t` to `s'`, `t'` and `s''`, `t''`, such that one of `(s', t')` and @@ -165,7 +165,7 @@ lemma cauchy_davenport_minOrder_mul (hs : s.Nonempty) (ht : t.Nonempty) : · rw [← card_smul_finset g⁻¹ t] refine Or.inr ((add_le_add_right hst _).trans ?_) rw [← card_union_of_disjoint hgt] - exact (card_le_card_mul_left _ hgs).trans (le_add_of_le_left aux1) + exact (card_le_card_mul_left hgs).trans (le_add_of_le_left aux1) -- Else, we're done by induction on either `(s', t')` or `(s'', t'')` depending on whether -- `|s| + |t| ≤ |s'| + |t'|` or `|s| + |t| < |s''| + |t''|`. One of those two inequalities must -- hold since `2 * (|s| + |t|) = |s'| + |t'| + |s''| + |t''|`. @@ -208,7 +208,7 @@ lemma cauchy_davenport_mul_of_linearOrder_isCancelMul [LinearOrder α] [Semigrou [MulLeftMono α] [MulRightMono α] {s t : Finset α} (hs : s.Nonempty) (ht : t.Nonempty) : #s + #t - 1 ≤ #(s * t) := by suffices s * {t.min' ht} ∩ ({s.max' hs} * t) = {s.max' hs * t.min' ht} by - rw [← card_singleton_mul t (s.max' hs), ← card_mul_singleton s (t.min' ht), + rw [← card_singleton_mul (s.max' hs) t, ← card_mul_singleton s (t.min' ht), ← card_union_add_card_inter, ← card_singleton _, ← this, Nat.add_sub_cancel] exact card_mono (union_subset (mul_subset_mul_left <| singleton_subset_iff.2 <| min'_mem _ _) <| mul_subset_mul_right <| singleton_subset_iff.2 <| max'_mem _ _) diff --git a/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean b/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean index 0a386abf7a0f9..fe073269bb474 100644 --- a/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean +++ b/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean @@ -179,7 +179,7 @@ theorem ruzsa_triangle_inequality_mul_mul_mul (A B C : Finset G) : push_cast refine (le_div_iff₀ <| cast_pos.2 hB.card_pos).1 ?_ rw [mul_div_right_comm, mul_comm _ B] - refine (Nat.cast_le.2 <| card_le_card_mul_left _ hU.1).trans ?_ + refine (Nat.cast_le.2 <| card_le_card_mul_left hU.1).trans ?_ refine le_trans ?_ (mul_le_mul (hUA _ hB') (cast_le.2 <| card_le_card <| mul_subset_mul_right hU.2) (zero_le _) (zero_le _)) From 698148763dd29cf98dd2c56b8a44ac8e825461ae Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 28 Nov 2024 20:56:22 +0000 Subject: [PATCH 404/829] =?UTF-8?q?feat:=20`stab=20s=20=E2=8A=93=20stab=20?= =?UTF-8?q?t=20=E2=89=A4=20stab=20(s=20=E2=88=A9=20t)`=20(#19477)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit and other similar lemmas Moves: * `mem_stabilizer_of_finite_iff_smul_le` -> `mem_stabilizer_set_iff_subset_smul_set` * `mem_stabilizer_of_finite_iff_le_smul` -> `mem_stabilizer_set_iff_smul_set_subset` From Kneser (LeanCamCombi) --- Mathlib/Algebra/Pointwise/Stabilizer.lean | 83 ++++++++++++++++++++- Mathlib/GroupTheory/GroupAction/Basic.lean | 26 ------- Mathlib/GroupTheory/GroupAction/Blocks.lean | 3 +- 3 files changed, 83 insertions(+), 29 deletions(-) diff --git a/Mathlib/Algebra/Pointwise/Stabilizer.lean b/Mathlib/Algebra/Pointwise/Stabilizer.lean index 0596f27478672..2a1b5ea298aad 100644 --- a/Mathlib/Algebra/Pointwise/Stabilizer.lean +++ b/Mathlib/Algebra/Pointwise/Stabilizer.lean @@ -16,11 +16,13 @@ open Function MulOpposite Set open scoped Pointwise namespace MulAction -variable {G H α : Type*} [Group G] [Group H] [MulAction G α] {a : G} +variable {G H α : Type*} /-! ### Stabilizer of a set -/ section Set +section Group +variable [Group G] [Group H] [MulAction G α] {a : G} {s t : Set α} @[to_additive (attr := simp)] lemma stabilizer_empty : stabilizer G (∅ : Set α) = ⊤ := @@ -57,8 +59,77 @@ lemma stabilizer_mul_self (s : Set G) : (stabilizer G s : Set G) * s = s := by rw [← mem_stabilizer_iff.1 ha] exact smul_mem_smul_set hb +@[to_additive] +lemma stabilizer_inf_stabilizer_le_stabilizer_apply₂ {f : Set α → Set α → Set α} + (hf : ∀ a : G, a • f s t = f (a • s) (a • t)) : + stabilizer G s ⊓ stabilizer G t ≤ stabilizer G (f s t) := by aesop (add simp [SetLike.le_def]) + +@[to_additive] +lemma stabilizer_inf_stabilizer_le_stabilizer_union : + stabilizer G s ⊓ stabilizer G t ≤ stabilizer G (s ∪ t) := + stabilizer_inf_stabilizer_le_stabilizer_apply₂ fun _ ↦ smul_set_union + +@[to_additive] +lemma stabilizer_inf_stabilizer_le_stabilizer_inter : + stabilizer G s ⊓ stabilizer G t ≤ stabilizer G (s ∩ t) := + stabilizer_inf_stabilizer_le_stabilizer_apply₂ fun _ ↦ smul_set_inter + +@[to_additive] +lemma stabilizer_inf_stabilizer_le_stabilizer_sdiff : + stabilizer G s ⊓ stabilizer G t ≤ stabilizer G (s \ t) := + stabilizer_inf_stabilizer_le_stabilizer_apply₂ fun _ ↦ smul_set_sdiff + +@[to_additive] +lemma stabilizer_union_eq_left (hdisj : Disjoint s t) (hstab : stabilizer G s ≤ stabilizer G t) + (hstab_union : stabilizer G (s ∪ t) ≤ stabilizer G t) : + stabilizer G (s ∪ t) = stabilizer G s := by + refine le_antisymm ?_ ?_ + · calc + stabilizer G (s ∪ t) + ≤ stabilizer G (s ∪ t) ⊓ stabilizer G t := by simpa + _ ≤ stabilizer G ((s ∪ t) \ t) := stabilizer_inf_stabilizer_le_stabilizer_sdiff + _ = stabilizer G s := by rw [union_diff_cancel_right]; simpa [← disjoint_iff_inter_eq_empty] + · calc + stabilizer G s + ≤ stabilizer G s ⊓ stabilizer G t := by simpa + _ ≤ stabilizer G (s ∪ t) := stabilizer_inf_stabilizer_le_stabilizer_union + +@[to_additive] +lemma stabilizer_union_eq_right (hdisj : Disjoint s t) (hstab : stabilizer G t ≤ stabilizer G s) + (hstab_union : stabilizer G (s ∪ t) ≤ stabilizer G s) : + stabilizer G (s ∪ t) = stabilizer G t := by + rw [union_comm, stabilizer_union_eq_left hdisj.symm hstab (union_comm .. ▸ hstab_union)] + +variable {s : Set G} + +open scoped RightActions in +@[to_additive] +lemma op_smul_set_stabilizer_subset (ha : a ∈ s) : (stabilizer G s : Set G) <• a ⊆ s := + smul_set_subset_iff.2 fun b hb ↦ by rw [← hb]; exact smul_mem_smul_set ha + +@[to_additive] +lemma stabilizer_subset_div_right (ha : a ∈ s) : ↑(stabilizer G s) ⊆ s / {a} := fun b hb ↦ + ⟨_, by rwa [← smul_eq_mul, mem_stabilizer_set.1 hb], _, mem_singleton _, mul_div_cancel_right _ _⟩ + +@[to_additive] +lemma stabilizer_finite (hs₀ : s.Nonempty) (hs : s.Finite) : (stabilizer G s : Set G).Finite := by + obtain ⟨a, ha⟩ := hs₀ + exact (hs.div <| finite_singleton _).subset <| stabilizer_subset_div_right ha + +end Group + +section CommGroup +variable [CommGroup G] {s t : Set G} {a : G} + +@[to_additive] +lemma smul_set_stabilizer_subset (ha : a ∈ s) : a • (stabilizer G s : Set G) ⊆ s := by + simpa using op_smul_set_stabilizer_subset ha + +end CommGroup end Set +variable [Group G] [Group H] [MulAction G α] {a : G} + /-! ### Stabilizer of a subgroup -/ section Subgroup @@ -92,7 +163,7 @@ end Subgroup section Finset variable [DecidableEq α] -@[to_additive (attr := simp)] +@[to_additive (attr := simp, norm_cast)] lemma stabilizer_coe_finset (s : Finset α) : stabilizer G (s : Set α) = stabilizer G s := by ext; simp [← Finset.coe_inj] @@ -132,6 +203,8 @@ end Finset /-! ### Stabilizer of a finite set -/ +variable {s : Set α} + @[to_additive] lemma mem_stabilizer_set_iff_subset_smul_set {s : Set α} (hs : s.Finite) : a ∈ stabilizer G s ↔ s ⊆ a • s := by @@ -148,6 +221,12 @@ lemma mem_stabilizer_set_iff_smul_set_subset {s : Set α} (hs : s.Finite) : rw [stabilizer_coe_finset, mem_stabilizer_finset_iff_smul_finset_subset, ← Finset.coe_smul_finset, Finset.coe_subset] +@[deprecated (since := "2024-11-25")] +alias mem_stabilizer_of_finite_iff_smul_le := mem_stabilizer_set_iff_subset_smul_set + +@[deprecated (since := "2024-11-25")] +alias mem_stabilizer_of_finite_iff_le_smul := mem_stabilizer_set_iff_smul_set_subset + @[to_additive] lemma mem_stabilizer_set' {s : Set α} (hs : s.Finite) : a ∈ stabilizer G s ↔ ∀ ⦃b⦄, b ∈ s → a • b ∈ s := by diff --git a/Mathlib/GroupTheory/GroupAction/Basic.lean b/Mathlib/GroupTheory/GroupAction/Basic.lean index ae7eaef28b78d..7ea1e10af8050 100644 --- a/Mathlib/GroupTheory/GroupAction/Basic.lean +++ b/Mathlib/GroupTheory/GroupAction/Basic.lean @@ -5,7 +5,6 @@ Authors: Chris Hughes -/ import Mathlib.Algebra.Group.Subgroup.Map import Mathlib.Data.Finite.Sigma -import Mathlib.Data.Set.Finite.Basic import Mathlib.Data.Set.Finite.Range import Mathlib.Data.Set.Pointwise.SMul import Mathlib.Data.Setoid.Basic @@ -291,29 +290,4 @@ theorem le_stabilizer_iff_smul_le (s : Set α) (H : Subgroup G) : simp only [Set.smul_mem_smul_set_iff, hx] · simp only [smul_inv_smul] -/-- To prove membership to stabilizer of a *finite set*, it is enough to prove one inclusion. -/ -theorem mem_stabilizer_of_finite_iff_smul_le (s : Set α) (hs : s.Finite) (g : G) : - g ∈ stabilizer G s ↔ g • s ⊆ s := by - haveI : Fintype s := Set.Finite.fintype hs - haveI : Finite (g • s : Set α) := Finite.Set.finite_image .. - haveI : Fintype (g • s : Set α) := Fintype.ofFinite _ - rw [mem_stabilizer_iff] - constructor - · exact Eq.subset - · rw [← Set.toFinset_inj, ← Set.toFinset_subset_toFinset] - intro h - apply Finset.eq_of_subset_of_card_le h - apply le_of_eq - suffices (g • s).toFinset = Finset.map ⟨_, MulAction.injective g⟩ hs.toFinset by - rw [this, Finset.card_map, Set.toFinite_toFinset] - rw [← Finset.coe_inj] - simp only [Set.coe_toFinset, Set.toFinite_toFinset, Finset.coe_map, - Function.Embedding.coeFn_mk, Set.image_smul] - -/-- To prove membership to stabilizer of a *finite set*, it is enough to prove one inclusion. -/ -theorem mem_stabilizer_of_finite_iff_le_smul (s : Set α) (hs : s.Finite) (g : G) : - g ∈ stabilizer G s ↔ s ⊆ g • s := by - rw [← @inv_mem_iff, mem_stabilizer_of_finite_iff_smul_le s hs] - exact Set.subset_set_smul_iff.symm - end MulAction diff --git a/Mathlib/GroupTheory/GroupAction/Blocks.lean b/Mathlib/GroupTheory/GroupAction/Blocks.lean index e3372fa4417f0..39dc66ecbf32e 100644 --- a/Mathlib/GroupTheory/GroupAction/Blocks.lean +++ b/Mathlib/GroupTheory/GroupAction/Blocks.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Antoine Chambert-Loir. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Antoine Chambert-Loir -/ +import Mathlib.Algebra.Pointwise.Stabilizer import Mathlib.Data.Setoid.Partition import Mathlib.GroupTheory.GroupAction.Pointwise import Mathlib.GroupTheory.GroupAction.SubMulAction @@ -553,7 +554,7 @@ theorem of_subset (a : X) (hfB : B.Finite) : smul_smul, ← mul_inv_rev] at hg hx ⊢ exact fun _ ↦ hx _ ∘ hg _ have hag' (g : G) (hg : a ∈ g • B') : B' = g • B' := by - rw [eq_comm, ← mem_stabilizer_iff, mem_stabilizer_of_finite_iff_le_smul _ hfB'] + rw [eq_comm, ← mem_stabilizer_iff, mem_stabilizer_set_iff_subset_smul_set hfB'] exact hag g hg rw [isBlock_iff_smul_eq_of_nonempty] rintro g ⟨b : X, hb' : b ∈ g • B', hb : b ∈ B'⟩ From d97300ddfac2e514d8daace021ca5b0fbca36866 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 29 Nov 2024 02:34:12 +0000 Subject: [PATCH 405/829] chore(Algebra/BigOperators/Group/Finset): golf some proofs (#19593) --- .../Algebra/BigOperators/Group/Finset.lean | 60 +++++++++---------- 1 file changed, 28 insertions(+), 32 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index 282177e1ba7c0..c235e6c2bd74a 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -1241,16 +1241,17 @@ lemma mulIndicator_prod (s : Finset ι) (t : Set κ) (f : ι → κ → β) : variable {κ : Type*} @[to_additive] -lemma mulIndicator_biUnion (s : Finset ι) (t : ι → Set κ) {f : κ → β} : - ((s : Set ι).PairwiseDisjoint t) → - mulIndicator (⋃ i ∈ s, t i) f = fun a ↦ ∏ i ∈ s, mulIndicator (t i) f a := by - classical - refine Finset.induction_on s (by simp) fun i s hi ih hs ↦ funext fun j ↦ ?_ - rw [prod_insert hi, set_biUnion_insert, mulIndicator_union_of_not_mem_inter, - ih (hs.subset <| subset_insert _ _)] - simp only [not_exists, exists_prop, mem_iUnion, mem_inter_iff, not_and] - exact fun hji i' hi' hji' ↦ (ne_of_mem_of_not_mem hi' hi).symm <| - hs.elim_set (mem_insert_self _ _) (mem_insert_of_mem hi') _ hji hji' +lemma mulIndicator_biUnion (s : Finset ι) (t : ι → Set κ) {f : κ → β} + (hs : (s : Set ι).PairwiseDisjoint t) : + mulIndicator (⋃ i ∈ s, t i) f = fun a ↦ ∏ i ∈ s, mulIndicator (t i) f a := by + induction s using Finset.cons_induction with + | empty => simp + | cons i s hi ih => + ext j + rw [coe_cons, Set.pairwiseDisjoint_insert_of_not_mem (Finset.mem_coe.not.2 hi)] at hs + classical + rw [prod_cons, cons_eq_insert, set_biUnion_insert, mulIndicator_union_of_not_mem_inter, ih hs.1] + exact (Set.disjoint_iff.mp (Set.disjoint_iUnion₂_right.mpr hs.2) ·) @[to_additive] lemma mulIndicator_biUnion_apply (s : Finset ι) (t : ι → Set κ) {f : κ → β} @@ -1734,13 +1735,12 @@ theorem prod_pow_boole [DecidableEq α] (s : Finset α) (f : α → β) (a : α) theorem prod_dvd_prod_of_dvd {S : Finset α} (g1 g2 : α → β) (h : ∀ a ∈ S, g1 a ∣ g2 a) : S.prod g1 ∣ S.prod g2 := by - classical - induction S using Finset.induction_on' with - | h₁ => simp - | h₂ _haS _hTS haT IH => - rename_i a T - rw [Finset.prod_insert haT, prod_insert haT] - exact mul_dvd_mul (h a <| T.mem_insert_self a) <| IH fun b hb ↦ h b <| mem_insert_of_mem hb + induction S using Finset.cons_induction with + | empty => simp + | cons a T haT IH => + rw [Finset.prod_cons, Finset.prod_cons] + rw [Finset.forall_mem_cons] at h + exact mul_dvd_mul h.1 <| IH h.2 theorem prod_dvd_prod_of_subset {ι M : Type*} [CommMonoid M] (s t : Finset ι) (f : ι → M) (h : s ⊆ t) : (∏ i ∈ s, f i) ∣ ∏ i ∈ t, f i := @@ -1887,10 +1887,9 @@ theorem card_eq_sum_card_image [DecidableEq β] (f : α → β) (s : Finset α) theorem mem_sum {f : α → Multiset β} (s : Finset α) (b : β) : (b ∈ ∑ x ∈ s, f x) ↔ ∃ a ∈ s, b ∈ f a := by - classical - refine s.induction_on (by simp) ?_ - intro a t hi ih - simp [sum_insert hi, ih, or_and_right, exists_or] + induction s using Finset.cons_induction with + | empty => simp + | cons a t hi ih => simp [sum_cons, ih, or_and_right, exists_or] @[to_additive] theorem prod_unique_nonempty {α β : Type*} [CommMonoid β] [Unique α] (s : Finset α) (f : α → β) @@ -2162,11 +2161,9 @@ theorem toFinset_prod_dvd_prod [CommMonoid α] (S : Multiset α) : S.toFinset.pr @[to_additive] theorem prod_sum {α : Type*} {ι : Type*} [CommMonoid α] (f : ι → Multiset α) (s : Finset ι) : (∑ x ∈ s, f x).prod = ∏ x ∈ s, (f x).prod := by - classical - induction s using Finset.induction_on with - | empty => rw [Finset.sum_empty, Finset.prod_empty, Multiset.prod_zero] - | insert hat ih => - rw [Finset.sum_insert hat, Finset.prod_insert hat, Multiset.prod_add, ih] + induction s using Finset.cons_induction with + | empty => rw [Finset.sum_empty, Finset.prod_empty, Multiset.prod_zero] + | cons a s has ih => rw [Finset.sum_cons, Finset.prod_cons, Multiset.prod_add, ih] end Multiset @@ -2177,12 +2174,11 @@ theorem Units.coe_prod {M : Type*} [CommMonoid M] (f : α → Mˣ) (s : Finset theorem nat_abs_sum_le {ι : Type*} (s : Finset ι) (f : ι → ℤ) : (∑ i ∈ s, f i).natAbs ≤ ∑ i ∈ s, (f i).natAbs := by - classical - induction s using Finset.induction_on with - | empty => simp only [Finset.sum_empty, Int.natAbs_zero, le_refl] - | insert his IH => - simp only [his, Finset.sum_insert, not_false_iff] - exact (Int.natAbs_add_le _ _).trans (Nat.add_le_add_left IH _) + induction s using Finset.cons_induction with + | empty => simp only [Finset.sum_empty, Int.natAbs_zero, le_refl] + | cons i s his IH => + simp only [Finset.sum_cons, not_false_iff] + exact (Int.natAbs_add_le _ _).trans (Nat.add_le_add_left IH _) /-! ### `Additive`, `Multiplicative` -/ From dade1b8db960d689139ae664cad3511b42c178d7 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 29 Nov 2024 03:02:44 +0000 Subject: [PATCH 406/829] chore: bump toolchain to v4.14.0-rc3 (#19594) This bumps the toolchain to v4.14.0-rc3. This differs from rc2 only in changes to the behaviour of Lake, which will hopefully eliminate problems with ProofWidgets. --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index 57a4710c03ea5..6d9e70f73356b 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.14.0-rc2 +leanprover/lean4:v4.14.0-rc3 From 27d7119d186eacf4da6b55bdb919e65560367fa2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Fri, 29 Nov 2024 03:48:28 +0000 Subject: [PATCH 407/829] chore(SetTheory/Ordinal/Basic): deprecate results on `WellOrder` (#18208) This PR is analogous to #18205. --- Mathlib/SetTheory/Ordinal/Basic.lean | 23 ++++++++++------------- 1 file changed, 10 insertions(+), 13 deletions(-) diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index ec9a34419fbf4..e7920e2f034c9 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -87,10 +87,8 @@ namespace WellOrder instance inhabited : Inhabited WellOrder := ⟨⟨PEmpty, _, inferInstanceAs (IsWellOrder PEmpty EmptyRelation)⟩⟩ -@[simp] -theorem eta (o : WellOrder) : mk o.α o.r o.wo = o := by - cases o - rfl +@[deprecated (since := "2024-10-24")] +theorem eta (o : WellOrder) : mk o.α o.r o.wo = o := rfl end WellOrder @@ -139,22 +137,21 @@ instance inhabited : Inhabited Ordinal := instance one : One Ordinal := ⟨type <| @EmptyRelation PUnit⟩ -@[simp] -theorem type_def' (w : WellOrder) : ⟦w⟧ = type w.r := by - cases w - rfl +/-- Avoid using `Quotient.mk` to construct an `Ordinal` directly -/ +@[deprecated (since := "2024-10-24")] +theorem type_def' (w : WellOrder) : ⟦w⟧ = type w.r := rfl -@[simp] -theorem type_def (r) [wo : IsWellOrder α r] : (⟦⟨α, r, wo⟩⟧ : Ordinal) = type r := by - rfl +/-- Avoid using `Quotient.mk` to construct an `Ordinal` directly -/ +@[deprecated (since := "2024-10-24")] +theorem type_def (r) [wo : IsWellOrder α r] : (⟦⟨α, r, wo⟩⟧ : Ordinal) = type r := rfl @[simp] theorem type_toType (o : Ordinal) : type (α := o.toType) (· < ·) = o := - (type_def' _).symm.trans <| Quotient.out_eq o + o.out_eq @[deprecated type_toType (since := "2024-10-22")] theorem type_lt (o : Ordinal) : type (α := o.toType) (· < ·) = o := - (type_def' _).symm.trans <| Quotient.out_eq o + o.out_eq @[deprecated type_toType (since := "2024-08-26")] theorem type_out (o : Ordinal) : Ordinal.type o.out.r = o := From 8e5beeaf96585866b1984b2364bc0f4ce9c29bb4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Fri, 29 Nov 2024 04:10:21 +0000 Subject: [PATCH 408/829] refactor(Order/WellFounded): deprecate `WellFounded.succ` (#18238) This definition is somewhat of a footgun, as it can't actually be used to define a `SuccOrder` instance on partial orders, and in the linear order case, one can use [`SuccOrder.ofLinearWellFoundedLT`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Basic.html#SuccOrder.ofLinearWellFoundedLT) for that purpose. --- Mathlib/Order/WellFounded.lean | 16 ++++++++++++---- 1 file changed, 12 insertions(+), 4 deletions(-) diff --git a/Mathlib/Order/WellFounded.lean b/Mathlib/Order/WellFounded.lean index cf572cbc31597..99a592719d4bc 100644 --- a/Mathlib/Order/WellFounded.lean +++ b/Mathlib/Order/WellFounded.lean @@ -88,21 +88,27 @@ protected theorem lt_sup {r : α → α → Prop} (wf : WellFounded r) {s : Set (hx : x ∈ s) : r x (wf.sup s h) := min_mem wf { x | ∀ a ∈ s, r a x } h x hx -section +section deprecated + +set_option linter.deprecated false open Classical in /-- A successor of an element `x` in a well-founded order is a minimal element `y` such that -`x < y` if one exists. Otherwise it is `x` itself. -/ +`x < y` if one exists. Otherwise it is `x` itself. + +Deprecated. If you have a linear order, consider defining a `SuccOrder` instance through +`ConditionallyCompleteLinearOrder.toSuccOrder`. -/ +@[deprecated (since := "2024-10-25")] protected noncomputable def succ {r : α → α → Prop} (wf : WellFounded r) (x : α) : α := if h : ∃ y, r x y then wf.min { y | r x y } h else x +@[deprecated (since := "2024-10-25")] protected theorem lt_succ {r : α → α → Prop} (wf : WellFounded r) {x : α} (h : ∃ y, r x y) : r x (wf.succ x) := by rw [WellFounded.succ, dif_pos h] apply min_mem -end - +@[deprecated (since := "2024-10-25")] protected theorem lt_succ_iff {r : α → α → Prop} [wo : IsWellOrder α r] {x : α} (h : ∃ y, r x y) (y : α) : r y (wo.wf.succ x) ↔ r y x ∨ y = x := by constructor @@ -120,6 +126,8 @@ protected theorem lt_succ_iff {r : α → α → Prop} [wo : IsWellOrder α r] { exact hy rintro (hy | rfl); (· exact _root_.trans hy (wo.wf.lt_succ h)); exact wo.wf.lt_succ h +end deprecated + end WellFounded section LinearOrder From a389037297bd67474802933f114704f6123f0fb5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Fri, 29 Nov 2024 04:10:22 +0000 Subject: [PATCH 409/829] feat(Order/InitialSeg): initial segments preserve successor limits (#19053) Initial segments preserve the covering relation, successors, and successor limits. --- Mathlib.lean | 1 + Mathlib/Order/InitialSeg.lean | 15 +++-- Mathlib/Order/SuccPred/Basic.lean | 2 +- Mathlib/Order/SuccPred/InitialSeg.lean | 79 ++++++++++++++++++++++++++ 4 files changed, 92 insertions(+), 5 deletions(-) create mode 100644 Mathlib/Order/SuccPred/InitialSeg.lean diff --git a/Mathlib.lean b/Mathlib.lean index ad4182f451904..a6e93d1a935f9 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4100,6 +4100,7 @@ import Mathlib.Order.Sublattice import Mathlib.Order.SuccPred.Archimedean import Mathlib.Order.SuccPred.Basic import Mathlib.Order.SuccPred.CompleteLinearOrder +import Mathlib.Order.SuccPred.InitialSeg import Mathlib.Order.SuccPred.IntervalSucc import Mathlib.Order.SuccPred.Limit import Mathlib.Order.SuccPred.LinearLocallyFinite diff --git a/Mathlib/Order/InitialSeg.lean b/Mathlib/Order/InitialSeg.lean index b09f50366e0c0..9d83ed92d1c08 100644 --- a/Mathlib/Order/InitialSeg.lean +++ b/Mathlib/Order/InitialSeg.lean @@ -7,6 +7,7 @@ import Mathlib.Data.Sum.Order import Mathlib.Logic.Equiv.Set import Mathlib.Order.RelIso.Set import Mathlib.Order.WellFounded + /-! # Initial and principal segments @@ -656,12 +657,15 @@ theorem monotone [PartialOrder α] (f : α ≤i β) : Monotone f := theorem strictMono [PartialOrder α] (f : α ≤i β) : StrictMono f := f.toOrderEmbedding.strictMono -theorem map_isMin [PartialOrder α] (f : α ≤i β) (h : IsMin a) : IsMin (f a) := by - intro b hb +@[simp] +theorem isMin_apply_iff [PartialOrder α] (f : α ≤i β) : IsMin (f a) ↔ IsMin a := by + refine ⟨StrictMono.isMin_of_apply f.strictMono, fun h b hb ↦ ?_⟩ obtain ⟨x, rfl⟩ := f.mem_range_of_le hb rw [f.le_iff_le] at hb ⊢ exact h hb +alias ⟨_, map_isMin⟩ := isMin_apply_iff + @[simp] theorem map_bot [PartialOrder α] [OrderBot α] [OrderBot β] (f : α ≤i β) : f ⊥ = ⊥ := (map_isMin f isMin_bot).eq_bot @@ -712,8 +716,11 @@ theorem monotone [PartialOrder α] (f : α intro h b hb + · rw [← f.apply_covBy_apply_iff] at hb + exact h _ hb + · obtain ⟨c, rfl⟩ := f.mem_range_of_rel hb.lt + rw [f.apply_covBy_apply_iff] at hb + exact h _ hb + +@[simp] +theorem isSuccLimit_apply_iff (f : α ≤i β) : IsSuccLimit (f a) ↔ IsSuccLimit a := by + simp [IsSuccLimit] + +end InitialSeg + +namespace PrincipalSeg + +@[simp] +theorem apply_covBy_apply_iff (f : α Date: Fri, 29 Nov 2024 05:06:30 +0000 Subject: [PATCH 410/829] chore: update Mathlib dependencies 2024-11-29 (#19595) This PR updates the Mathlib dependencies. --- lake-manifest.json | 62 +++++++++++++++++++++++----------------------- 1 file changed, 31 insertions(+), 31 deletions(-) diff --git a/lake-manifest.json b/lake-manifest.json index 57203830de6b6..e799426a1c291 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -1,34 +1,34 @@ {"version": "1.1.0", "packagesDir": ".lake/packages", "packages": - [{"url": "https://github.com/leanprover-community/batteries", + [{"url": "https://github.com/leanprover-community/plausible", "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "c933dd9b00271d869e22b802a015092d1e8e454a", - "name": "batteries", + "rev": "42dc02bdbc5d0c2f395718462a76c3d87318f7fa", + "name": "plausible", "manifestFile": "lake-manifest.json", "inputRev": "main", "inherited": false, "configFile": "lakefile.toml"}, - {"url": "https://github.com/leanprover-community/quote4", + {"url": "https://github.com/leanprover-community/LeanSearchClient", "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "303b23fbcea94ac4f96e590c1cad6618fd4f5f41", - "name": "Qq", + "rev": "d7caecce0d0f003fd5e9cce9a61f1dd6ba83142b", + "name": "LeanSearchClient", "manifestFile": "lake-manifest.json", - "inputRev": "master", + "inputRev": "main", "inherited": false, - "configFile": "lakefile.lean"}, - {"url": "https://github.com/leanprover-community/aesop", + "configFile": "lakefile.toml"}, + {"url": "https://github.com/leanprover-community/import-graph", "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "de91b59101763419997026c35a41432ac8691f15", - "name": "aesop", + "rev": "119b022b3ea88ec810a677888528e50f8144a26e", + "name": "importGraph", "manifestFile": "lake-manifest.json", - "inputRev": "master", + "inputRev": "main", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover-community/ProofWidgets4", @@ -41,45 +41,45 @@ "inputRev": "v0.0.46", "inherited": false, "configFile": "lakefile.lean"}, - {"url": "https://github.com/leanprover/lean4-cli", + {"url": "https://github.com/leanprover-community/aesop", "type": "git", "subDir": null, - "scope": "leanprover", - "rev": "726b3c9ad13acca724d4651f14afc4804a7b0e4d", - "name": "Cli", + "scope": "leanprover-community", + "rev": "de91b59101763419997026c35a41432ac8691f15", + "name": "aesop", "manifestFile": "lake-manifest.json", - "inputRev": "main", - "inherited": true, + "inputRev": "master", + "inherited": false, "configFile": "lakefile.toml"}, - {"url": "https://github.com/leanprover-community/import-graph", + {"url": "https://github.com/leanprover-community/quote4", "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "119b022b3ea88ec810a677888528e50f8144a26e", - "name": "importGraph", + "rev": "303b23fbcea94ac4f96e590c1cad6618fd4f5f41", + "name": "Qq", "manifestFile": "lake-manifest.json", - "inputRev": "main", + "inputRev": "master", "inherited": false, - "configFile": "lakefile.toml"}, - {"url": "https://github.com/leanprover-community/LeanSearchClient", + "configFile": "lakefile.lean"}, + {"url": "https://github.com/leanprover-community/batteries", "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "d7caecce0d0f003fd5e9cce9a61f1dd6ba83142b", - "name": "LeanSearchClient", + "rev": "c933dd9b00271d869e22b802a015092d1e8e454a", + "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", "inherited": false, "configFile": "lakefile.toml"}, - {"url": "https://github.com/leanprover-community/plausible", + {"url": "https://github.com/leanprover/lean4-cli", "type": "git", "subDir": null, - "scope": "leanprover-community", - "rev": "42dc02bdbc5d0c2f395718462a76c3d87318f7fa", - "name": "plausible", + "scope": "leanprover", + "rev": "726b3c9ad13acca724d4651f14afc4804a7b0e4d", + "name": "Cli", "manifestFile": "lake-manifest.json", "inputRev": "main", - "inherited": false, + "inherited": true, "configFile": "lakefile.toml"}], "name": "mathlib", "lakeDir": ".lake"} From ab780650a6c8db2942d137b91ede7c94c47a4950 Mon Sep 17 00:00:00 2001 From: damiano Date: Fri, 29 Nov 2024 06:22:59 +0000 Subject: [PATCH 411/829] fix: add `realm_emoji` (and `emoji_code`) to auto-zulip reactions (#19597) [Zulip discussion](https://leanprover.zulipchat.com/#narrow/channel/144837-PR-reviews/topic/.2319522.3A.20emoji.20reactions.20to.20all.20public.20channels/near/485003969) --- scripts/zulip_emoji_merge_delegate.py | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/scripts/zulip_emoji_merge_delegate.py b/scripts/zulip_emoji_merge_delegate.py index f31e66c7b506f..f734f4f48d351 100644 --- a/scripts/zulip_emoji_merge_delegate.py +++ b/scripts/zulip_emoji_merge_delegate.py @@ -82,7 +82,9 @@ print('Removing bors') result = client.remove_reaction({ "message_id": message['id'], - "emoji_name": "bors" + "emoji_name": "bors", + "emoji_code": "22134", + "reaction_type": "realm_emoji", }) print(f"result: '{result}'") if has_merge: From 1426b9e52291eec21fbe16e7f2b2757ab2ca26f1 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Fri, 29 Nov 2024 08:51:12 +0000 Subject: [PATCH 412/829] chore: update Mathlib dependencies 2024-11-29 (#19598) 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 e799426a1c291..42df1c2b614e6 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "c933dd9b00271d869e22b802a015092d1e8e454a", + "rev": "dfbe9387054e2972449de7bf346059d3609529fa", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From 8f06fcac7450cf3e5fc78bcbcb53acfd74121828 Mon Sep 17 00:00:00 2001 From: Pim Otte Date: Fri, 29 Nov 2024 09:51:36 +0000 Subject: [PATCH 413/829] feat(Combinatorics/SimpleGraph/Matching): add `IsCycles` and `IsAlternating` with basic results (#19250) add `SimpleGraph.IsCycles` and `SimpleGraph.IsAlternating` with some basic results involving matchings. These are defined on `SimpleGraph`, because we want to use the `symmDiff` defined on there. This is in preparation for Tutte's theorem. Co-authored-by: Pim Otte --- .../Combinatorics/SimpleGraph/Matching.lean | 87 +++++++++++++++++++ 1 file changed, 87 insertions(+) diff --git a/Mathlib/Combinatorics/SimpleGraph/Matching.lean b/Mathlib/Combinatorics/SimpleGraph/Matching.lean index d52240d416332..5d6ee34bb110b 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Matching.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Matching.lean @@ -27,6 +27,14 @@ one edge, and the edges of the subgraph represent the paired vertices. * `SimpleGraph.Subgraph.IsPerfectMatching` defines when a subgraph `M` of a simple graph is a perfect matching, denoted `M.IsPerfectMatching`. +* `SimpleGraph.IsMatchingFree` means that a graph `G` has no perfect matchings. + +* `SimpleGraph.IsCycles` means that a graph consists of cycles (including cycles of length 0, + also known as isolated vertices) + +* `SimpleGraph.IsAlternating` means that edges in a graph `G` are alternatingly + included and not included in some other graph `G'` + ## TODO * Define an `other` function and prove useful results about it (https://leanprover.zulipchat.com/#narrow/stream/252551-graph-theory/topic/matchings/near/266205863) @@ -309,4 +317,83 @@ lemma exists_maximal_isMatchingFree [Finite V] (h : G.IsMatchingFree) : obtain ⟨Gmax, hGmax⟩ := Finite.exists_le_maximal h exact ⟨Gmax, ⟨hGmax.1, ⟨hGmax.2.prop, fun _ h' ↦ hGmax.2.not_prop_of_gt h'⟩⟩⟩ +/-- A graph `G` consists of a set of cycles, if each vertex is either isolated or connected to +exactly two vertices. This is used to create new matchings by taking the `symmDiff` with cycles. +The definition of `symmDiff` that makes sense is the one for `SimpleGraph`. The `symmDiff` +for `SimpleGraph.Subgraph` deriving from the lattice structure also affects the vertices included, +which we do not want in this case. This is why this property is defined for `SimpleGraph`, rather +than `SimpleGraph.Subgraph`. +-/ +def IsCycles (G : SimpleGraph V) := ∀ ⦃v⦄, (G.neighborSet v).Nonempty → (G.neighborSet v).ncard = 2 + +/-- +Given a vertex with one edge in a graph of cycles this gives the other edge incident +to the same vertex. +-/ +lemma IsCycles.other_adj_of_adj (h : G.IsCycles) (hadj : G.Adj v w) : + ∃ w', w ≠ w' ∧ G.Adj v w' := by + simp_rw [← SimpleGraph.mem_neighborSet] at hadj ⊢ + cases' (G.neighborSet v).eq_empty_or_nonempty with hl hr + · exact ((Set.eq_empty_iff_forall_not_mem.mp hl) _ hadj).elim + · have := h hr + obtain ⟨w', hww'⟩ := (G.neighborSet v).exists_ne_of_one_lt_ncard (by omega) w + exact ⟨w', ⟨hww'.2.symm, hww'.1⟩⟩ + +open scoped symmDiff + +lemma Subgraph.IsPerfectMatching.symmDiff_spanningCoe_IsCycles + {M : Subgraph G} {M' : Subgraph G'} (hM : M.IsPerfectMatching) + (hM' : M'.IsPerfectMatching) : (M.spanningCoe ∆ M'.spanningCoe).IsCycles := by + intro v + obtain ⟨w, hw⟩ := hM.1 (hM.2 v) + obtain ⟨w', hw'⟩ := hM'.1 (hM'.2 v) + simp only [symmDiff_def, Set.ncard_eq_two, ne_eq, imp_iff_not_or, Set.not_nonempty_iff_eq_empty, + Set.eq_empty_iff_forall_not_mem, SimpleGraph.mem_neighborSet, SimpleGraph.sup_adj, sdiff_adj, + spanningCoe_adj, not_or, not_and, not_not] + by_cases hww' : w = w' + · simp_all [← imp_iff_not_or, hww'] + · right + use w, w' + aesop + +/-- +A graph `G` is alternating with respect to some other graph `G'`, if exactly every other edge in +`G` is in `G'`. Note that the degree of each vertex needs to be at most 2 for this to be +possible. This property is used to create new matchings using `symmDiff`. +The definition of `symmDiff` that makes sense is the one for `SimpleGraph`. The `symmDiff` +for `SimpleGraph.Subgraph` deriving from the lattice structure also affects the vertices included, +which we do not want in this case. This is why this property, just like `IsCycles`, is defined +for `SimpleGraph` rather than `SimpleGraph.Subgraph`. +-/ +def IsAlternating (G G' : SimpleGraph V) := + ∀ ⦃v w w': V⦄, w ≠ w' → G.Adj v w → G.Adj v w' → (G'.Adj v w ↔ ¬ G'.Adj v w') + +lemma IsPerfectMatching.symmDiff_spanningCoe_of_isAlternating {M : Subgraph G} + (hM : M.IsPerfectMatching) (hG' : G'.IsAlternating M.spanningCoe) (hG'cyc : G'.IsCycles) : + (SimpleGraph.toSubgraph (M.spanningCoe ∆ G') + (by rfl)).IsPerfectMatching := by + rw [Subgraph.isPerfectMatching_iff] + intro v + simp only [toSubgraph_adj, symmDiff_def, sup_adj, sdiff_adj, Subgraph.spanningCoe_adj] + obtain ⟨w, hw⟩ := hM.1 (hM.2 v) + by_cases h : G'.Adj v w + · obtain ⟨w', hw'⟩ := hG'cyc.other_adj_of_adj h + have hmadj : M.Adj v w ↔ ¬M.Adj v w' := by simpa using hG' hw'.1 h hw'.2 + use w' + simp only [hmadj.mp hw.1, hw'.2, not_true_eq_false, and_self, not_false_eq_true, or_true, + true_and] + rintro y (hl | hr) + · aesop + · obtain ⟨w'', hw''⟩ := hG'cyc.other_adj_of_adj hr.1 + by_contra! hc + simp_all only [show M.Adj v y ↔ ¬M.Adj v w' from by simpa using hG' hc hr.1 hw'.2, + not_false_eq_true, ne_eq, iff_true, not_true_eq_false, and_false] + · use w + simp only [hw.1, h, not_false_eq_true, and_self, not_true_eq_false, or_false, true_and] + rintro y (hl | hr) + · exact hw.2 _ hl.1 + · have ⟨w', hw'⟩ := hG'cyc.other_adj_of_adj hr.1 + simp_all only [show M.Adj v y ↔ ¬M.Adj v w' from by simpa using hG' hw'.1 hr.1 hw'.2, not_not, + ne_eq, and_false] + end SimpleGraph From 9c9b448e95a8fad15d2b4c79b79c7e3c9a11caf4 Mon Sep 17 00:00:00 2001 From: blizzard_inc Date: Fri, 29 Nov 2024 10:59:57 +0000 Subject: [PATCH 414/829] feat: Label zulip PR links with awaiting-author emoji (#19599) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add action to add/remove the ✍️ emoji reaction to links on Zulip to PRs that get the `awaiting-review` label. Co-authored-by: blizzard_inc --- .../zulip_emoji_awaiting_author.yaml | 34 +++++++++++++++++++ scripts/zulip_emoji_merge_delegate.py | 18 ++++++++++ 2 files changed, 52 insertions(+) create mode 100644 .github/workflows/zulip_emoji_awaiting_author.yaml diff --git a/.github/workflows/zulip_emoji_awaiting_author.yaml b/.github/workflows/zulip_emoji_awaiting_author.yaml new file mode 100644 index 0000000000000..e3f57ab9cfc82 --- /dev/null +++ b/.github/workflows/zulip_emoji_awaiting_author.yaml @@ -0,0 +1,34 @@ +on: + pull_request: + types: [labeled, unlabeled] +jobs: + set_pr_emoji: + if: github.event.label.name == 'awaiting-author' + runs-on: ubuntu-latest + steps: + - name: Checkout mathlib repository + uses: actions/checkout@v4 + with: + sparse-checkout: | + scripts/zulip_emoji_merge_delegate.py + + - name: Set up Python + uses: actions/setup-python@v5 + with: + python-version: '3.x' + + - name: Install dependencies + run: | + python -m pip install --upgrade pip + pip install zulip + + - name: Add or remove emoji + env: + ZULIP_API_KEY: ${{ secrets.ZULIP_API_KEY }} + ZULIP_EMAIL: github-mathlib4-bot@leanprover.zulipchat.com + ZULIP_SITE: https://leanprover.zulipchat.com + PR_NUMBER: ${{ github.event.number}} + LABEL_STATUS: ${{ github.event.action }} + run: | + printf $'Running the python script with pr "%s"\n' "$PR_NUMBER" + python scripts/zulip_emoji_merge_delegate.py "$ZULIP_API_KEY" "$ZULIP_EMAIL" "$ZULIP_SITE" "$LABEL_STATUS" "$PR_NUMBER" diff --git a/scripts/zulip_emoji_merge_delegate.py b/scripts/zulip_emoji_merge_delegate.py index f734f4f48d351..de3e5d3ca70c3 100644 --- a/scripts/zulip_emoji_merge_delegate.py +++ b/scripts/zulip_emoji_merge_delegate.py @@ -65,6 +65,7 @@ has_peace_sign = any(reaction['emoji_name'] == 'peace_sign' for reaction in reactions) has_bors = any(reaction['emoji_name'] == 'bors' for reaction in reactions) has_merge = any(reaction['emoji_name'] == 'merge' for reaction in reactions) + has_awaiting_author = any(reaction['emoji_name'] == 'writing' for reaction in reactions) match = pr_pattern.search(content) if match: print(f"matched: '{message}'") @@ -94,6 +95,14 @@ "emoji_name": "merge" }) print(f"result: '{result}'") + if has_awaiting_author: + print('Removing awaiting-author') + result = client.remove_reaction({ + "message_id": message['id'], + "emoji_name": "writing" + }) + print(f"result: '{result}'") + # applying appropriate emoji reaction print("Applying reactions, as appropriate.") @@ -109,6 +118,15 @@ "message_id": message['id'], "emoji_name": "peace_sign" }) + elif LABEL == 'labeled': + print('adding awaiting-author') + client.add_reaction({ + "message_id": message['id'], + "emoji_name": "writing" + }) + elif LABEL == 'unlabeled': + print('awaiting-author removed') + # the reaction was already removed. elif LABEL.startswith("[Merged by Bors]"): print('adding [Merged by Bors]') client.add_reaction({ From 90f062feea36adad53b81925aa9dc3e09c268262 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Fri, 29 Nov 2024 11:09:14 +0000 Subject: [PATCH 415/829] feat(NumberTheory/EllipticDivisibilitySequence): extend even-odd recursion to integers (#13786) Provide a quick proof that normalised EDSs satisfy an even-odd recursion over all integers rather than just natural numbers, without relying on heavier machinery i.e. the full EDS recurrence in #13155, and make `Int.negInduction` slightly stronger. --- .../DivisionPolynomial/Basic.lean | 330 ++++++++++-------- .../DivisionPolynomial/Degree.lean | 69 ++-- Mathlib/Data/Int/Defs.lean | 4 +- Mathlib/Data/Nat/Cast/Defs.lean | 4 + .../EllipticDivisibilitySequence.lean | 183 +++++++--- 5 files changed, 360 insertions(+), 230 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean index 15a2fccfcd2b3..d5026ebed2afd 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean @@ -15,75 +15,70 @@ These are defined in terms of the auxiliary sequences for normalised elliptic di ## Mathematical background -Let $W$ be a Weierstrass curve over a commutative ring $R$. The sequence of $n$-division polynomials -$\psi_n \in R[X, Y]$ of $W$ is the normalised EDS with initial values - * $\psi_0 := 0$, - * $\psi_1 := 1$, - * $\psi_2 := 2Y + a_1X + a_3$, - * $\psi_3 := 3X^4 + b_2X^3 + 3b_4X ^ 2 + 3b_6X + b_8$, and - * $\psi_4 := \psi_2 \cdot - (2X^6 + b_2X^5 + 5b_4X^4 + 10b_6X^3 + 10b_8X^2 + (b_2b_8 - b_4b_6)X + (b_4b_8 - b_6^2))$. - -Furthermore, define the associated sequences $\phi_n, \omega_n \in R[X, Y]$ by - * $\phi_n := X\psi_n^2 - \psi_{n + 1} \cdot \psi_{n - 1}$, and - * $\omega_n := (\psi_{2n} / \psi_n - \psi_n \cdot (a_1\phi_n + a_3\psi_n^2)) / 2$. - -Note that $\omega_n$ is always well-defined as a polynomial in $R[X, Y]$. As a start, it can be -shown by induction that $\psi_n$ always divides $\psi_{2n}$ in $R[X, Y]$, so that -$\psi_{2n} / \psi_n$ is always well-defined as a polynomial, while division by 2 is well-defined -when $R$ has characteristic different from 2. In general, it can be shown that 2 always divides the -polynomial $\psi_{2n} / \psi_n - \psi_n \cdot (a_1\phi_n + a_3\psi_n^2)$ in the characteristic zero -universal ring $\mathcal{R}[X, Y] := \mathbb{Z}[A_1, A_2, A_3, A_4, A_6][X, Y]$ of $W$, where the -$A_i$ are indeterminates. Then $\omega_n$ can be equivalently defined as the image of this division -under the associated universal morphism $\mathcal{R}[X, Y] \to R[X, Y]$ mapping $A_i$ to $a_i$. - -Now, in the coordinate ring $R[W]$, note that $\psi_2^2$ is congruent to the polynomial -$\Psi_2^{[2]} := 4X^3 + b_2X^2 + 2b_4X + b_6 \in R[X]$. As such, in $R[W]$, the recurrences -associated to a normalised EDS show that $\psi_n / \psi_2$ are congruent to certain polynomials in -$R[X]$. In particular, define $\tilde{\Psi}_n \in R[X]$ as the auxiliary sequence for a normalised -EDS with extra parameter $(\Psi_2^{[2]})^2$ and initial values - * $\tilde{\Psi}_0 := 0$, - * $\tilde{\Psi}_1 := 1$, - * $\tilde{\Psi}_2 := 1$, - * $\tilde{\Psi}_3 := \psi_3$, and - * $\tilde{\Psi}_4 := \psi_4 / \psi_2$. - -The corresponding normalised EDS $\Psi_n \in R[X, Y]$ is then given by - * $\Psi_n := \tilde{\Psi}_n \cdot \psi_2$ if $n$ is even, and - * $\Psi_n := \tilde{\Psi}_n$ if $n$ is odd. - -Furthermore, define the associated sequences $\Psi_n^{[2]}, \Phi_n \in R[X]$ by - * $\Psi_n^{[2]} := \tilde{\Psi}_n^2 \cdot \Psi_2^{[2]}$ if $n$ is even, - * $\Psi_n^{[2]} := \tilde{\Psi}_n^2$ if $n$ is odd, - * $\Phi_n := X\Psi_n^{[2]} - \tilde{\Psi}_{n + 1} \cdot \tilde{\Psi}_{n - 1}$ if $n$ is even, and - * $\Phi_n := X\Psi_n^{[2]} - \tilde{\Psi}_{n + 1} \cdot \tilde{\Psi}_{n - 1} \cdot \Psi_2^{[2]}$ - if $n$ is odd. - -With these definitions in mind, $\psi_n \in R[X, Y]$ and $\phi_n \in R[X, Y]$ are congruent in -$R[W]$ to $\Psi_n \in R[X, Y]$ and $\Phi_n \in R[X]$ respectively, which are defined in terms of -$\Psi_2^{[2]} \in R[X]$ and $\tilde{\Psi}_n \in R[X]$. +Let `W` be a Weierstrass curve over a commutative ring `R`. The sequence of `n`-division polynomials +`ψₙ ∈ R[X, Y]` of `W` is the normalised EDS with initial values + * `ψ₀ := 0`, + * `ψ₁ := 1`, + * `ψ₂ := 2Y + a₁X + a₃`, + * `ψ₃ := 3X⁴ + b₂X³ + 3b₄X² + 3b₆X + b₈`, and + * `ψ₄ := ψ₂ ⬝ (2X⁶ + b₂X⁵ + 5b₄X⁴ + 10b₆X³ + 10b₈X² + (b₂b₈ - b₄b₆)X + (b₄b₈ - b₆²))`. + +Furthermore, define the associated sequences `φₙ, ωₙ ∈ R[X, Y]` by + * `φₙ := Xψₙ² - ψₙ₊₁ ⬝ ψₙ₋₁`, and + * `ωₙ := (ψ₂ₙ / ψₙ - ψₙ ⬝ (a₁φₙ + a₃ψₙ²)) / 2`. + +Note that `ωₙ` is always well-defined as a polynomial in `R[X, Y]`. As a start, it can be shown by +induction that `ψₙ` always divides `ψ₂ₙ` in `R[X, Y]`, so that `ψ₂ₙ / ψₙ` is always well-defined as +a polynomial, while division by `2` is well-defined when `R` has characteristic different from `2`. +In general, it can be shown that `2` always divides the polynomial `ψ₂ₙ / ψₙ - ψₙ ⬝ (a₁φₙ + a₃ψₙ²)` +in the characteristic `0` universal ring `𝓡[X, Y] := ℤ[A₁, A₂, A₃, A₄, A₆][X, Y]` of `W`, where the +`Aᵢ` are indeterminates. Then `ωₙ` can be equivalently defined as the image of this division under +the associated universal morphism `𝓡[X, Y] → R[X, Y]` mapping `Aᵢ` to `aᵢ`. + +Now, in the coordinate ring `R[W]`, note that `ψ₂²` is congruent to the polynomial +`Ψ₂Sq := 4X³ + b₂X² + 2b₄X + b₆ ∈ R[X]`. As such, the recurrences of a normalised EDS show that +`ψₙ / ψ₂` are congruent to certain polynomials in `R[W]`. In particular, define `preΨₙ ∈ R[X]` as +the auxiliary sequence for a normalised EDS with extra parameter `Ψ₂Sq²` and initial values + * `preΨ₀ := 0`, + * `preΨ₁ := 1`, + * `preΨ₂ := 1`, + * `preΨ₃ := ψ₃`, and + * `preΨ₄ := ψ₄ / ψ₂`. + +The corresponding normalised EDS `Ψₙ ∈ R[X, Y]` is then given by + * `Ψₙ := preΨₙ ⬝ ψ₂` if `n` is even, and + * `Ψₙ := preΨₙ` if `n` is odd. + +Furthermore, define the associated sequences `ΨSqₙ, Φₙ ∈ R[X]` by + * `ΨSqₙ := preΨₙ² ⬝ Ψ₂Sq` if `n` is even, + * `ΨSqₙ := preΨₙ²` if `n` is odd, + * `Φₙ := XΨSqₙ - preΨₙ₊₁ ⬝ preΨₙ₋₁` if `n` is even, and + * `Φₙ := XΨSqₙ - preΨₙ₊₁ ⬝ preΨₙ₋₁ ⬝ Ψ₂Sq` if `n` is odd. + +With these definitions, `ψₙ ∈ R[X, Y]` and `φₙ ∈ R[X, Y]` are congruent in `R[W]` to `Ψₙ ∈ R[X, Y]` +and `Φₙ ∈ R[X]` respectively, which are defined in terms of `Ψ₂Sq ∈ R[X]` and `preΨₙ ∈ R[X]`. ## Main definitions - * `WeierstrassCurve.preΨ`: the univariate polynomials $\tilde{\Psi}_n$. - * `WeierstrassCurve.ΨSq`: the univariate polynomials $\Psi_n^{[2]}$. - * `WeierstrassCurve.Ψ`: the bivariate polynomials $\Psi_n$. - * `WeierstrassCurve.Φ`: the univariate polynomials $\Phi_n$. - * `WeierstrassCurve.ψ`: the bivariate $n$-division polynomials $\psi_n$. - * `WeierstrassCurve.φ`: the bivariate polynomials $\phi_n$. - * TODO: the bivariate polynomials $\omega_n$. + * `WeierstrassCurve.preΨ`: the univariate polynomials `preΨₙ`. + * `WeierstrassCurve.ΨSq`: the univariate polynomials `ΨSqₙ`. + * `WeierstrassCurve.Ψ`: the bivariate polynomials `Ψₙ`. + * `WeierstrassCurve.Φ`: the univariate polynomials `Φₙ`. + * `WeierstrassCurve.ψ`: the bivariate `n`-division polynomials `ψₙ`. + * `WeierstrassCurve.φ`: the bivariate polynomials `φₙ`. + * TODO: the bivariate polynomials `ωₙ`. ## Implementation notes Analogously to `Mathlib.NumberTheory.EllipticDivisibilitySequence`, the bivariate polynomials -$\Psi_n$ are defined in terms of the univariate polynomials $\tilde{\Psi}_n$. This is done partially -to avoid ring division, but more crucially to allow the definition of $\Psi_n^{[2]}$ and $\Phi_n$ as -univariate polynomials without needing to work under the coordinate ring, and to allow the -computation of their leading terms without ambiguity. Furthermore, evaluating these polynomials at a -rational point on $W$ recovers their original definition up to linear combinations of the -Weierstrass equation of $W$, hence also avoiding the need to work in the coordinate ring. +`Ψₙ` are defined in terms of the univariate polynomials `preΨₙ`. This is done partially to avoid +ring division, but more crucially to allow the definition of `ΨSqₙ` and `Φₙ` as univariate +polynomials without needing to work under the coordinate ring, and to allow the computation of their +leading terms without ambiguity. Furthermore, evaluating these polynomials at a rational point on +`W` recovers their original definition up to linear combinations of the Weierstrass equation of `W`, +hence also avoiding the need to work in the coordinate ring. -TODO: implementation notes for the definition of $\omega_n$. +TODO: implementation notes for the definition of `ωₙ`. ## References @@ -114,13 +109,13 @@ variable {R : Type r} {S : Type s} [CommRing R] [CommRing S] (W : WeierstrassCur section Ψ₂Sq -/-! ### The univariate polynomial $\Psi_2^{[2]}$ -/ +/-! ### The univariate polynomial `Ψ₂Sq` -/ -/-- The $2$-division polynomial $\psi_2 = \Psi_2$. -/ +/-- The `2`-division polynomial `ψ₂ = Ψ₂`. -/ noncomputable def ψ₂ : R[X][Y] := W.toAffine.polynomialY -/-- The univariate polynomial $\Psi_2^{[2]}$ congruent to $\psi_2^2$. -/ +/-- The univariate polynomial `Ψ₂Sq` congruent to `ψ₂²`. -/ noncomputable def Ψ₂Sq : R[X] := C 4 * X ^ 3 + C W.b₂ * X ^ 2 + C (2 * W.b₄) * X + C W.b₆ @@ -143,20 +138,20 @@ end Ψ₂Sq section preΨ' -/-! ### The univariate polynomials $\tilde{\Psi}_n$ for $n \in \mathbb{N}$ -/ +/-! ### The univariate polynomials `preΨₙ` for `n ∈ ℕ` -/ -/-- The $3$-division polynomial $\psi_3 = \Psi_3$. -/ +/-- The `3`-division polynomial `ψ₃ = Ψ₃`. -/ noncomputable def Ψ₃ : R[X] := 3 * X ^ 4 + C W.b₂ * X ^ 3 + 3 * C W.b₄ * X ^ 2 + 3 * C W.b₆ * X + C W.b₈ -/-- The univariate polynomial $\tilde{\Psi}_4$, which is auxiliary to the $4$-division polynomial -$\psi_4 = \Psi_4 = \tilde{\Psi}_4\psi_2$. -/ +/-- The univariate polynomial `preΨ₄`, which is auxiliary to the 4-division polynomial +`ψ₄ = Ψ₄ = preΨ₄ψ₂`. -/ noncomputable def preΨ₄ : R[X] := 2 * X ^ 6 + C W.b₂ * X ^ 5 + 5 * C W.b₄ * X ^ 4 + 10 * C W.b₆ * X ^ 3 + 10 * C W.b₈ * X ^ 2 + C (W.b₂ * W.b₈ - W.b₄ * W.b₆) * X + C (W.b₄ * W.b₈ - W.b₆ ^ 2) -/-- The univariate polynomials $\tilde{\Psi}_n$ for $n \in \mathbb{N}$, which are auxiliary to the -bivariate polynomials $\Psi_n$ congruent to the bivariate $n$-division polynomials $\psi_n$. -/ +/-- The univariate polynomials `preΨₙ` for `n ∈ ℕ`, which are auxiliary to the bivariate polynomials +`Ψₙ` congruent to the bivariate `n`-division polynomials `ψₙ`. -/ noncomputable def preΨ' (n : ℕ) : R[X] := preNormEDS' (W.Ψ₂Sq ^ 2) W.Ψ₃ W.preΨ₄ n @@ -180,24 +175,24 @@ lemma preΨ'_three : W.preΨ' 3 = W.Ψ₃ := lemma preΨ'_four : W.preΨ' 4 = W.preΨ₄ := preNormEDS'_four .. -lemma preΨ'_odd (m : ℕ) : W.preΨ' (2 * (m + 2) + 1) = - W.preΨ' (m + 4) * W.preΨ' (m + 2) ^ 3 * (if Even m then W.Ψ₂Sq ^ 2 else 1) - - W.preΨ' (m + 1) * W.preΨ' (m + 3) ^ 3 * (if Even m then 1 else W.Ψ₂Sq ^ 2) := - preNormEDS'_odd .. - lemma preΨ'_even (m : ℕ) : W.preΨ' (2 * (m + 3)) = W.preΨ' (m + 2) ^ 2 * W.preΨ' (m + 3) * W.preΨ' (m + 5) - W.preΨ' (m + 1) * W.preΨ' (m + 3) * W.preΨ' (m + 4) ^ 2 := preNormEDS'_even .. +lemma preΨ'_odd (m : ℕ) : W.preΨ' (2 * (m + 2) + 1) = + W.preΨ' (m + 4) * W.preΨ' (m + 2) ^ 3 * (if Even m then W.Ψ₂Sq ^ 2 else 1) - + W.preΨ' (m + 1) * W.preΨ' (m + 3) ^ 3 * (if Even m then 1 else W.Ψ₂Sq ^ 2) := + preNormEDS'_odd .. + end preΨ' section preΨ -/-! ### The univariate polynomials $\tilde{\Psi}_n$ for $n \in \mathbb{Z}$ -/ +/-! ### The univariate polynomials `preΨₙ` for `n ∈ ℤ` -/ -/-- The univariate polynomials $\tilde{\Psi}_n$ for $n \in \mathbb{Z}$, which are auxiliary to the -bivariate polynomials $\Psi_n$ congruent to the bivariate $n$-division polynomials $\psi_n$. -/ +/-- The univariate polynomials `preΨₙ` for `n ∈ ℤ`, which are auxiliary to the bivariate polynomials +`Ψₙ` congruent to the bivariate `n`-division polynomials `ψₙ`. -/ noncomputable def preΨ (n : ℤ) : R[X] := preNormEDS (W.Ψ₂Sq ^ 2) W.Ψ₃ W.preΨ₄ n @@ -225,27 +220,37 @@ lemma preΨ_three : W.preΨ 3 = W.Ψ₃ := lemma preΨ_four : W.preΨ 4 = W.preΨ₄ := preNormEDS_four .. -lemma preΨ_odd (m : ℕ) : W.preΨ (2 * (m + 2) + 1) = - W.preΨ (m + 4) * W.preΨ (m + 2) ^ 3 * (if Even m then W.Ψ₂Sq ^ 2 else 1) - - W.preΨ (m + 1) * W.preΨ (m + 3) ^ 3 * (if Even m then 1 else W.Ψ₂Sq ^ 2) := - preNormEDS_odd .. - -lemma preΨ_even (m : ℕ) : W.preΨ (2 * (m + 3)) = +lemma preΨ_even_ofNat (m : ℕ) : W.preΨ (2 * (m + 3)) = W.preΨ (m + 2) ^ 2 * W.preΨ (m + 3) * W.preΨ (m + 5) - W.preΨ (m + 1) * W.preΨ (m + 3) * W.preΨ (m + 4) ^ 2 := - preNormEDS_even .. + preNormEDS_even_ofNat .. + +lemma preΨ_odd_ofNat (m : ℕ) : W.preΨ (2 * (m + 2) + 1) = + W.preΨ (m + 4) * W.preΨ (m + 2) ^ 3 * (if Even m then W.Ψ₂Sq ^ 2 else 1) - + W.preΨ (m + 1) * W.preΨ (m + 3) ^ 3 * (if Even m then 1 else W.Ψ₂Sq ^ 2) := + preNormEDS_odd_ofNat .. @[simp] lemma preΨ_neg (n : ℤ) : W.preΨ (-n) = -W.preΨ n := preNormEDS_neg .. +lemma preΨ_even (m : ℤ) : W.preΨ (2 * m) = + W.preΨ (m - 1) ^ 2 * W.preΨ m * W.preΨ (m + 2) - + W.preΨ (m - 2) * W.preΨ m * W.preΨ (m + 1) ^ 2 := + preNormEDS_even .. + +lemma preΨ_odd (m : ℤ) : W.preΨ (2 * m + 1) = + W.preΨ (m + 2) * W.preΨ m ^ 3 * (if Even m then W.Ψ₂Sq ^ 2 else 1) - + W.preΨ (m - 1) * W.preΨ (m + 1) ^ 3 * (if Even m then 1 else W.Ψ₂Sq ^ 2) := + preNormEDS_odd .. + end preΨ section ΨSq -/-! ### The univariate polynomials $\Psi_n^{[2]}$ -/ +/-! ### The univariate polynomials `ΨSqₙ` -/ -/-- The univariate polynomials $\Psi_n^{[2]}$ congruent to $\psi_n^2$. -/ +/-- The univariate polynomials `ΨSqₙ` congruent to `ψₙ²`. -/ noncomputable def ΨSq (n : ℤ) : R[X] := W.preΨ n ^ 2 * if Even n then W.Ψ₂Sq else 1 @@ -255,45 +260,55 @@ lemma ΨSq_ofNat (n : ℕ) : W.ΨSq n = W.preΨ' n ^ 2 * if Even n then W.Ψ₂S @[simp] lemma ΨSq_zero : W.ΨSq 0 = 0 := by - erw [ΨSq_ofNat, preΨ'_zero, zero_pow two_ne_zero, zero_mul] + rw [← Nat.cast_zero, ΨSq_ofNat, preΨ'_zero, zero_pow two_ne_zero, zero_mul] @[simp] lemma ΨSq_one : W.ΨSq 1 = 1 := by - erw [ΨSq_ofNat, preΨ'_one, one_pow, mul_one] + rw [← Nat.cast_one, ΨSq_ofNat, preΨ'_one, one_pow, one_mul, if_neg Nat.not_even_one] @[simp] lemma ΨSq_two : W.ΨSq 2 = W.Ψ₂Sq := by - erw [ΨSq_ofNat, preΨ'_two, one_pow, one_mul, if_pos even_two] + rw [← Nat.cast_two, ΨSq_ofNat, preΨ'_two, one_pow, one_mul, if_pos even_two] @[simp] lemma ΨSq_three : W.ΨSq 3 = W.Ψ₃ ^ 2 := by - erw [ΨSq_ofNat, preΨ'_three, mul_one] + rw [← Nat.cast_three, ΨSq_ofNat, preΨ'_three, if_neg <| by decide, mul_one] @[simp] lemma ΨSq_four : W.ΨSq 4 = W.preΨ₄ ^ 2 * W.Ψ₂Sq := by - erw [ΨSq_ofNat, preΨ'_four, if_pos <| by decide] - -lemma ΨSq_odd (m : ℕ) : W.ΨSq (2 * (m + 2) + 1) = - (W.preΨ' (m + 4) * W.preΨ' (m + 2) ^ 3 * (if Even m then W.Ψ₂Sq ^ 2 else 1) - - W.preΨ' (m + 1) * W.preΨ' (m + 3) ^ 3 * (if Even m then 1 else W.Ψ₂Sq ^ 2)) ^ 2 := by - erw [ΨSq_ofNat, preΨ'_odd, if_neg (m + 2).not_even_two_mul_add_one, mul_one] + rw [← Nat.cast_four, ΨSq_ofNat, preΨ'_four, if_pos <| by decide] -lemma ΨSq_even (m : ℕ) : W.ΨSq (2 * (m + 3)) = +lemma ΨSq_even_ofNat (m : ℕ) : W.ΨSq (2 * (m + 3)) = (W.preΨ' (m + 2) ^ 2 * W.preΨ' (m + 3) * W.preΨ' (m + 5) - W.preΨ' (m + 1) * W.preΨ' (m + 3) * W.preΨ' (m + 4) ^ 2) ^ 2 * W.Ψ₂Sq := by - erw [ΨSq_ofNat, preΨ'_even, if_pos <| even_two_mul _] + rw_mod_cast [ΨSq_ofNat, preΨ'_even, if_pos <| even_two_mul _] + +lemma ΨSq_odd_ofNat (m : ℕ) : W.ΨSq (2 * (m + 2) + 1) = + (W.preΨ' (m + 4) * W.preΨ' (m + 2) ^ 3 * (if Even m then W.Ψ₂Sq ^ 2 else 1) - + W.preΨ' (m + 1) * W.preΨ' (m + 3) ^ 3 * (if Even m then 1 else W.Ψ₂Sq ^ 2)) ^ 2 := by + rw_mod_cast [ΨSq_ofNat, preΨ'_odd, if_neg (m + 2).not_even_two_mul_add_one, mul_one] @[simp] lemma ΨSq_neg (n : ℤ) : W.ΨSq (-n) = W.ΨSq n := by simp only [ΨSq, preΨ_neg, neg_sq, even_neg] +lemma ΨSq_even (m : ℤ) : W.ΨSq (2 * m) = + (W.preΨ (m - 1) ^ 2 * W.preΨ m * W.preΨ (m + 2) - + W.preΨ (m - 2) * W.preΨ m * W.preΨ (m + 1) ^ 2) ^ 2 * W.Ψ₂Sq := by + rw [ΨSq, preΨ_even, if_pos <| even_two_mul _] + +lemma ΨSq_odd (m : ℤ) : W.ΨSq (2 * m + 1) = + (W.preΨ (m + 2) * W.preΨ m ^ 3 * (if Even m then W.Ψ₂Sq ^ 2 else 1) - + W.preΨ (m - 1) * W.preΨ (m + 1) ^ 3 * (if Even m then 1 else W.Ψ₂Sq ^ 2)) ^ 2 := by + rw [ΨSq, preΨ_odd, if_neg m.not_even_two_mul_add_one, mul_one] + end ΨSq section Ψ -/-! ### The bivariate polynomials $\Psi_n$ -/ +/-! ### The bivariate polynomials `Ψₙ` -/ -/-- The bivariate polynomials $\Psi_n$ congruent to the $n$-division polynomials $\psi_n$. -/ +/-- The bivariate polynomials `Ψₙ` congruent to the `n`-division polynomials `ψₙ`. -/ protected noncomputable def Ψ (n : ℤ) : R[X][Y] := C (W.preΨ n) * if Even n then W.ψ₂ else 1 @@ -305,43 +320,60 @@ lemma Ψ_ofNat (n : ℕ) : W.Ψ n = C (W.preΨ' n) * if Even n then W.ψ₂ else @[simp] lemma Ψ_zero : W.Ψ 0 = 0 := by - erw [Ψ_ofNat, preΨ'_zero, C_0, zero_mul] + rw [← Nat.cast_zero, Ψ_ofNat, preΨ'_zero, C_0, zero_mul] @[simp] lemma Ψ_one : W.Ψ 1 = 1 := by - erw [Ψ_ofNat, preΨ'_one, C_1, mul_one] + rw [← Nat.cast_one, Ψ_ofNat, preΨ'_one, C_1, if_neg Nat.not_even_one, mul_one] @[simp] lemma Ψ_two : W.Ψ 2 = W.ψ₂ := by - erw [Ψ_ofNat, preΨ'_two, one_mul, if_pos even_two] + rw [← Nat.cast_two, Ψ_ofNat, preΨ'_two, C_1, one_mul, if_pos even_two] @[simp] lemma Ψ_three : W.Ψ 3 = C W.Ψ₃ := by - erw [Ψ_ofNat, preΨ'_three, mul_one] + rw [← Nat.cast_three, Ψ_ofNat, preΨ'_three, if_neg <| by decide, mul_one] @[simp] lemma Ψ_four : W.Ψ 4 = C W.preΨ₄ * W.ψ₂ := by - erw [Ψ_ofNat, preΨ'_four, if_pos <| by decide] + rw [← Nat.cast_four, Ψ_ofNat, preΨ'_four, if_pos <| by decide] -lemma Ψ_odd (m : ℕ) : W.Ψ (2 * (m + 2) + 1) = - W.Ψ (m + 4) * W.Ψ (m + 2) ^ 3 - W.Ψ (m + 1) * W.Ψ (m + 3) ^ 3 + - W.toAffine.polynomial * (16 * W.toAffine.polynomial - 8 * W.ψ₂ ^ 2) * C - (if Even m then W.preΨ' (m + 4) * W.preΨ' (m + 2) ^ 3 - else -W.preΨ' (m + 1) * W.preΨ' (m + 3) ^ 3) := by - repeat erw [Ψ_ofNat] - simp_rw [preΨ'_odd, if_neg (m + 2).not_even_two_mul_add_one, Nat.even_add_one, ite_not] - split_ifs <;> C_simp <;> rw [C_Ψ₂Sq] <;> ring1 - -lemma Ψ_even (m : ℕ) : W.Ψ (2 * (m + 3)) * W.ψ₂ = +lemma Ψ_even_ofNat (m : ℕ) : W.Ψ (2 * (m + 3)) * W.ψ₂ = W.Ψ (m + 2) ^ 2 * W.Ψ (m + 3) * W.Ψ (m + 5) - W.Ψ (m + 1) * W.Ψ (m + 3) * W.Ψ (m + 4) ^ 2 := by - repeat erw [Ψ_ofNat] + repeat rw_mod_cast [Ψ_ofNat] simp_rw [preΨ'_even, if_pos <| even_two_mul _, Nat.even_add_one, ite_not] split_ifs <;> C_simp <;> ring1 +lemma Ψ_odd_ofNat (m : ℕ) : W.Ψ (2 * (m + 2) + 1) = + W.Ψ (m + 4) * W.Ψ (m + 2) ^ 3 - W.Ψ (m + 1) * W.Ψ (m + 3) ^ 3 + + W.toAffine.polynomial * (16 * W.toAffine.polynomial - 8 * W.ψ₂ ^ 2) * + C (if Even m then W.preΨ' (m + 4) * W.preΨ' (m + 2) ^ 3 + else -W.preΨ' (m + 1) * W.preΨ' (m + 3) ^ 3) := by + repeat rw_mod_cast [Ψ_ofNat] + simp_rw [preΨ'_odd, if_neg (m + 2).not_even_two_mul_add_one, Nat.even_add_one, ite_not] + split_ifs <;> C_simp <;> rw [C_Ψ₂Sq] <;> ring1 + @[simp] lemma Ψ_neg (n : ℤ) : W.Ψ (-n) = -W.Ψ n := by simp only [Ψ, preΨ_neg, C_neg, neg_mul (α := R[X][Y]), even_neg] +lemma Ψ_even (m : ℤ) : W.Ψ (2 * m) * W.ψ₂ = + W.Ψ (m - 1) ^ 2 * W.Ψ m * W.Ψ (m + 2) - W.Ψ (m - 2) * W.Ψ m * W.Ψ (m + 1) ^ 2 := by + repeat rw [Ψ] + simp_rw [preΨ_even, if_pos <| even_two_mul _, Int.even_add_one, show m + 2 = m + 1 + 1 by ring1, + Int.even_add_one, show m - 2 = m - 1 - 1 by ring1, Int.even_sub_one, ite_not] + split_ifs <;> C_simp <;> ring1 + +lemma Ψ_odd (m : ℤ) : W.Ψ (2 * m + 1) = + W.Ψ (m + 2) * W.Ψ m ^ 3 - W.Ψ (m - 1) * W.Ψ (m + 1) ^ 3 + + W.toAffine.polynomial * (16 * W.toAffine.polynomial - 8 * W.ψ₂ ^ 2) * + C (if Even m then W.preΨ (m + 2) * W.preΨ m ^ 3 + else -W.preΨ (m - 1) * W.preΨ (m + 1) ^ 3) := by + repeat rw [Ψ] + simp_rw [preΨ_odd, if_neg m.not_even_two_mul_add_one, show m + 2 = m + 1 + 1 by ring1, + Int.even_add_one, Int.even_sub_one, ite_not] + split_ifs <;> C_simp <;> rw [C_Ψ₂Sq] <;> ring1 + lemma Affine.CoordinateRing.mk_Ψ_sq (n : ℤ) : mk W (W.Ψ n) ^ 2 = mk W (C <| W.ΨSq n) := by simp only [Ψ, ΨSq, map_one, map_mul, map_pow, one_pow, mul_pow, ite_pow, apply_ite C, apply_ite <| mk W, mk_ψ₂_sq] @@ -350,9 +382,9 @@ end Ψ section Φ -/-! ### The univariate polynomials $\Phi_n$ -/ +/-! ### The univariate polynomials `Φₙ` -/ -/-- The univariate polynomials $\Phi_n$ congruent to $\phi_n$. -/ +/-- The univariate polynomials `Φₙ` congruent to `φₙ`. -/ protected noncomputable def Φ (n : ℤ) : R[X] := X * W.ΨSq n - W.preΨ (n + 1) * W.preΨ (n - 1) * if Even n then 1 else W.Ψ₂Sq @@ -362,7 +394,8 @@ open WeierstrassCurve (Φ) lemma Φ_ofNat (n : ℕ) : W.Φ (n + 1) = X * W.preΨ' (n + 1) ^ 2 * (if Even n then 1 else W.Ψ₂Sq) - W.preΨ' (n + 2) * W.preΨ' n * (if Even n then W.Ψ₂Sq else 1) := by - erw [Φ, ΨSq_ofNat, ← mul_assoc, preΨ_ofNat, add_sub_cancel_right, preΨ_ofNat] + rw [Φ, ← Nat.cast_one, ← Nat.cast_add, ΨSq_ofNat, ← mul_assoc, ← Nat.cast_add, preΨ_ofNat, + Nat.cast_add, add_sub_cancel_right, preΨ_ofNat, ← Nat.cast_add] simp only [Nat.even_add_one, Int.even_add_one, Int.even_coe_nat, ite_not] @[simp] @@ -372,22 +405,26 @@ lemma Φ_zero : W.Φ 0 = 1 := by @[simp] lemma Φ_one : W.Φ 1 = X := by - erw [Φ_ofNat, preΨ'_one, one_pow, mul_one, mul_one, preΨ'_zero, mul_zero, zero_mul, sub_zero] + rw [show 1 = ((0 : ℕ) + 1 : ℤ) by rfl, Φ_ofNat, preΨ'_one, one_pow, mul_one, if_pos even_zero, + mul_one, preΨ'_zero, mul_zero, zero_mul, sub_zero] @[simp] lemma Φ_two : W.Φ 2 = X ^ 4 - C W.b₄ * X ^ 2 - C (2 * W.b₆) * X - C W.b₈ := by - erw [Φ_ofNat, preΨ'_two, if_neg Nat.not_even_one, Ψ₂Sq, preΨ'_three, preΨ'_one, mul_one, Ψ₃] + rw [show 2 = ((1 : ℕ) + 1 : ℤ) by rfl, Φ_ofNat, preΨ'_two, if_neg Nat.not_even_one, Ψ₂Sq, + preΨ'_three, preΨ'_one, if_neg Nat.not_even_one, Ψ₃] C_simp ring1 @[simp] lemma Φ_three : W.Φ 3 = X * W.Ψ₃ ^ 2 - W.preΨ₄ * W.Ψ₂Sq := by - erw [Φ_ofNat, preΨ'_three, mul_one, preΨ'_four, preΨ'_two, mul_one, if_pos even_two] + rw [show 3 = ((2 : ℕ) + 1 : ℤ) by rfl, Φ_ofNat, preΨ'_three, if_pos <| by decide, mul_one, + preΨ'_four, preΨ'_two, mul_one, if_pos even_two] @[simp] lemma Φ_four : W.Φ 4 = X * W.preΨ₄ ^ 2 * W.Ψ₂Sq - W.Ψ₃ * (W.preΨ₄ * W.Ψ₂Sq ^ 2 - W.Ψ₃ ^ 3) := by - erw [Φ_ofNat, preΨ'_four, if_neg <| by decide, show 3 + 2 = 2 * 2 + 1 by rfl, preΨ'_odd, - preΨ'_four, preΨ'_two, if_pos even_zero, preΨ'_one, preΨ'_three, if_pos even_zero] + rw [show 4 = ((3 : ℕ) + 1 : ℤ) by rfl, Φ_ofNat, preΨ'_four, if_neg <| by decide, + show 3 + 2 = 2 * 2 + 1 by rfl, preΨ'_odd, preΨ'_four, preΨ'_two, if_pos even_zero, preΨ'_one, + preΨ'_three, if_pos even_zero, if_neg <| by decide] ring1 @[simp] @@ -399,9 +436,9 @@ end Φ section ψ -/-! ### The bivariate polynomials $\psi_n$ -/ +/-! ### The bivariate polynomials `ψₙ` -/ -/-- The bivariate $n$-division polynomials $\psi_n$. -/ +/-- The bivariate `n`-division polynomials `ψₙ`. -/ protected noncomputable def ψ (n : ℤ) : R[X][Y] := normEDS W.ψ₂ (C W.Ψ₃) (C W.preΨ₄) n @@ -427,18 +464,26 @@ lemma ψ_three : W.ψ 3 = C W.Ψ₃ := lemma ψ_four : W.ψ 4 = C W.preΨ₄ * W.ψ₂ := normEDS_four .. -lemma ψ_odd (m : ℕ) : W.ψ (2 * (m + 2) + 1) = - W.ψ (m + 4) * W.ψ (m + 2) ^ 3 - W.ψ (m + 1) * W.ψ (m + 3) ^ 3 := - normEDS_odd .. - -lemma ψ_even (m : ℕ) : W.ψ (2 * (m + 3)) * W.ψ₂ = +lemma ψ_even_ofNat (m : ℕ) : W.ψ (2 * (m + 3)) * W.ψ₂ = W.ψ (m + 2) ^ 2 * W.ψ (m + 3) * W.ψ (m + 5) - W.ψ (m + 1) * W.ψ (m + 3) * W.ψ (m + 4) ^ 2 := - normEDS_even .. + normEDS_even_ofNat .. + +lemma ψ_odd_ofNat (m : ℕ) : W.ψ (2 * (m + 2) + 1) = + W.ψ (m + 4) * W.ψ (m + 2) ^ 3 - W.ψ (m + 1) * W.ψ (m + 3) ^ 3 := + normEDS_odd_ofNat .. @[simp] lemma ψ_neg (n : ℤ) : W.ψ (-n) = -W.ψ n := normEDS_neg .. +lemma ψ_even (m : ℤ) : W.ψ (2 * m) * W.ψ₂ = + W.ψ (m - 1) ^ 2 * W.ψ m * W.ψ (m + 2) - W.ψ (m - 2) * W.ψ m * W.ψ (m + 1) ^ 2 := + normEDS_even .. + +lemma ψ_odd (m : ℤ) : W.ψ (2 * m + 1) = + W.ψ (m + 2) * W.ψ m ^ 3 - W.ψ (m - 1) * W.ψ (m + 1) ^ 3 := + normEDS_odd .. + lemma Affine.CoordinateRing.mk_ψ (n : ℤ) : mk W (W.ψ n) = mk W (W.Ψ n) := by simp only [ψ, normEDS, Ψ, preΨ, map_mul, map_pow, map_preNormEDS, ← mk_ψ₂_sq, ← pow_mul] @@ -446,9 +491,9 @@ end ψ section φ -/-! ### The bivariate polynomials $\phi_n$ -/ +/-! ### The bivariate polynomials `φₙ` -/ -/-- The bivariate polynomials $\phi_n$. -/ +/-- The bivariate polynomials `φₙ`. -/ protected noncomputable def φ (n : ℤ) : R[X][Y] := C X * W.ψ n ^ 2 - W.ψ (n + 1) * W.ψ (n - 1) @@ -456,25 +501,28 @@ open WeierstrassCurve (Ψ Φ φ) @[simp] lemma φ_zero : W.φ 0 = 1 := by - erw [φ, ψ_zero, zero_pow two_ne_zero, mul_zero, zero_sub, ψ_one, one_mul, - zero_sub, ψ_neg, neg_neg, ψ_one] + rw [φ, ψ_zero, zero_pow two_ne_zero, mul_zero, zero_sub, zero_add, ψ_one, one_mul, zero_sub, + ψ_neg, neg_neg, ψ_one] @[simp] lemma φ_one : W.φ 1 = C X := by - erw [φ, ψ_one, one_pow, mul_one, ψ_zero, mul_zero, sub_zero] + rw [φ, ψ_one, one_pow, mul_one, sub_self, ψ_zero, mul_zero, sub_zero] @[simp] lemma φ_two : W.φ 2 = C X * W.ψ₂ ^ 2 - C W.Ψ₃ := by - erw [φ, ψ_two, ψ_three, ψ_one, mul_one] + rw [φ, ψ_two, two_add_one_eq_three, ψ_three, show (2 - 1 : ℤ) = 1 by rfl, ψ_one, mul_one] @[simp] lemma φ_three : W.φ 3 = C X * C W.Ψ₃ ^ 2 - C W.preΨ₄ * W.ψ₂ ^ 2 := by - erw [φ, ψ_three, ψ_four, mul_assoc, ψ_two, ← sq] + rw [φ, ψ_three, three_add_one_eq_four, ψ_four, mul_assoc, show (3 - 1 : ℤ) = 2 by rfl, ψ_two, + ← sq] @[simp] lemma φ_four : W.φ 4 = C X * C W.preΨ₄ ^ 2 * W.ψ₂ ^ 2 - C W.preΨ₄ * W.ψ₂ ^ 4 * C W.Ψ₃ + C W.Ψ₃ ^ 4 := by - erw [φ, ψ_four, show (4 + 1 : ℤ) = 2 * 2 + 1 by rfl, ψ_odd, ψ_four, ψ_two, ψ_one, ψ_three] + rw [φ, ψ_four, show (4 + 1 : ℤ) = 2 * 2 + 1 by rfl, ψ_odd, two_add_two_eq_four, ψ_four, + show (2 - 1 : ℤ) = 1 by rfl, ψ_two, ψ_one, two_add_one_eq_three, show (4 - 1 : ℤ) = 3 by rfl, + ψ_three] ring1 @[simp] diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Degree.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Degree.lean index e8193f45e38bd..ac58cdee1d2ca 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Degree.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Degree.lean @@ -14,29 +14,29 @@ Weierstrass curves defined in `Mathlib.AlgebraicGeometry.EllipticCurve.DivisionP ## Mathematical background -Let $W$ be a Weierstrass curve over a commutative ring $R$. By strong induction, - * $\tilde{\Psi}_n = \tfrac{n}{2}X^{\tfrac{n^2 - 4}{2}} + \dots$ if $n$ is even, - * $\tilde{\Psi}_n = nX^{\tfrac{n^2 - 1}{2}} + \dots$ if $n$ is odd, - * $\Psi_n^{[2]} = n^2X^{n^2 - 1} + \dots$, and - * $\Phi_n = X^{n^2} + \dots$. +Let `W` be a Weierstrass curve over a commutative ring `R`. By strong induction, + * `preΨₙ` has leading coefficient `n / 2` and degree `(n² - 4) / 2` if `n` is even, + * `preΨₙ` has leading coefficient `n` and degree `(n² - 1) / 2` if `n` is odd, + * `ΨSqₙ` has leading coefficient `n²` and degree `n² - 1`, and + * `Φₙ` has leading coefficient `1` and degree `n²`. -In particular, when $R$ is an integral domain of characteristic different from $n$, the univariate -polynomials $\tilde{\Psi}_n$, $\Psi_n^{[2]}$, and $\Phi_n$ all have their expected leading terms. +In particular, when `R` is an integral domain of characteristic different from `n`, the univariate +polynomials `preΨₙ`, `ΨSqₙ`, and `Φₙ` all have their expected leading terms. ## Main statements - * `WeierstrassCurve.natDegree_preΨ_le`: the expected degree of $\tilde{\Psi}_n$. - * `WeierstrassCurve.coeff_preΨ`: the expected leading coefficient of $\tilde{\Psi}_n$. - * `WeierstrassCurve.natDegree_preΨ`: the degree of $\tilde{\Psi}_n$ when $n \ne 0$. - * `WeierstrassCurve.leadingCoeff_preΨ`: the leading coefficient of $\tilde{\Psi}_n$ when $n \ne 0$. - * `WeierstrassCurve.natDegree_ΨSq_le`: the expected degree of $\Psi_n^{[2]}$. - * `WeierstrassCurve.coeff_ΨSq`: the expected leading coefficient of $\Psi_n^{[2]}$. - * `WeierstrassCurve.natDegree_ΨSq`: the degree of $\Psi_n^{[2]}$ when $n \ne 0$. - * `WeierstrassCurve.leadingCoeff_ΨSq`: the leading coefficient of $\Psi_n^{[2]}$ when $n \ne 0$. - * `WeierstrassCurve.natDegree_Φ_le`: the expected degree of $\Phi_n$. - * `WeierstrassCurve.coeff_Φ`: the expected leading coefficient of $\Phi_n$. - * `WeierstrassCurve.natDegree_Φ`: the degree of $\Phi_n$ when $n \ne 0$. - * `WeierstrassCurve.leadingCoeff_Φ`: the leading coefficient of $\Phi_n$ when $n \ne 0$. + * `WeierstrassCurve.natDegree_preΨ_le`: the degree bound `d` of `preΨₙ`. + * `WeierstrassCurve.coeff_preΨ`: the `d`-th coefficient of `preΨₙ`. + * `WeierstrassCurve.natDegree_preΨ`: the degree of `preΨₙ` when `n ≠ 0`. + * `WeierstrassCurve.leadingCoeff_preΨ`: the leading coefficient of `preΨₙ` when `n ≠ 0`. + * `WeierstrassCurve.natDegree_ΨSq_le`: the degree bound `d` of `ΨSqₙ`. + * `WeierstrassCurve.coeff_ΨSq`: the `d`-th coefficient of `ΨSqₙ`. + * `WeierstrassCurve.natDegree_ΨSq`: the degree of `ΨSqₙ` when `n ≠ 0`. + * `WeierstrassCurve.leadingCoeff_ΨSq`: the leading coefficient of `ΨSqₙ` when `n ≠ 0`. + * `WeierstrassCurve.natDegree_Φ_le`: the degree bound `d` of `Φₙ`. + * `WeierstrassCurve.coeff_Φ`: the `d`-th coefficient of `Φₙ`. + * `WeierstrassCurve.natDegree_Φ`: the degree of `Φₙ` when `n ≠ 0`. + * `WeierstrassCurve.leadingCoeff_Φ`: the leading coefficient of `Φₙ` when `n ≠ 0`. ## References @@ -271,16 +271,16 @@ lemma natDegree_preΨ_le (n : ℤ) : (W.preΨ n).natDegree ≤ (n.natAbs ^ 2 - if Even n then 4 else 1) / 2 := by induction n using Int.negInduction with | nat n => exact_mod_cast W.preΨ_ofNat n ▸ W.natDegree_preΨ'_le n - | neg => simpa only [preΨ_neg, natDegree_neg, Int.natAbs_neg, even_neg] + | neg ih => simp only [preΨ_neg, natDegree_neg, Int.natAbs_neg, even_neg, ih] @[simp] lemma coeff_preΨ (n : ℤ) : (W.preΨ n).coeff ((n.natAbs ^ 2 - if Even n then 4 else 1) / 2) = if Even n then n / 2 else n := by induction n using Int.negInduction with | nat n => exact_mod_cast W.preΨ_ofNat n ▸ W.coeff_preΨ' n - | neg n ih => + | neg ih n => simp only [preΨ_neg, coeff_neg, Int.natAbs_neg, even_neg] - rcases n.even_or_odd' with ⟨n, rfl | rfl⟩ <;> + rcases ih n, n.even_or_odd' with ⟨ih, ⟨n, rfl | rfl⟩⟩ <;> push_cast [even_two_mul, Int.not_even_two_mul_add_one, Int.neg_ediv_of_dvd ⟨n, rfl⟩] at * <;> rw [ih] @@ -289,8 +289,8 @@ lemma coeff_preΨ_ne_zero {n : ℤ} (h : (n : R) ≠ 0) : induction n using Int.negInduction with | nat n => simpa only [preΨ_ofNat, Int.even_coe_nat] using W.coeff_preΨ'_ne_zero <| by exact_mod_cast h - | neg n ih => simpa only [preΨ_neg, coeff_neg, neg_ne_zero, Int.natAbs_neg, even_neg] - using ih <| neg_ne_zero.mp <| by exact_mod_cast h + | neg ih n => simpa only [preΨ_neg, coeff_neg, neg_ne_zero, Int.natAbs_neg, even_neg] + using ih n <| neg_ne_zero.mp <| by exact_mod_cast h @[simp] lemma natDegree_preΨ {n : ℤ} (h : (n : R) ≠ 0) : @@ -301,8 +301,8 @@ lemma natDegree_preΨ_pos {n : ℤ} (hn : 2 < n.natAbs) (h : (n : R) ≠ 0) : 0 < (W.preΨ n).natDegree := by induction n using Int.negInduction with | nat n => simpa only [preΨ_ofNat] using W.natDegree_preΨ'_pos hn <| by exact_mod_cast h - | neg n ih => simpa only [preΨ_neg, natDegree_neg] - using ih (by rwa [← Int.natAbs_neg]) <| neg_ne_zero.mp <| by exact_mod_cast h + | neg ih n => simpa only [preΨ_neg, natDegree_neg] + using ih n (by rwa [← Int.natAbs_neg]) <| neg_ne_zero.mp <| by exact_mod_cast h @[simp] lemma leadingCoeff_preΨ {n : ℤ} (h : (n : R) ≠ 0) : @@ -312,7 +312,8 @@ lemma leadingCoeff_preΨ {n : ℤ} (h : (n : R) ≠ 0) : lemma preΨ_ne_zero [Nontrivial R] {n : ℤ} (h : (n : R) ≠ 0) : W.preΨ n ≠ 0 := by induction n using Int.negInduction with | nat n => simpa only [preΨ_ofNat] using W.preΨ'_ne_zero <| by exact_mod_cast h - | neg n ih => simpa only [preΨ_neg, neg_ne_zero] using ih <| neg_ne_zero.mp <| by exact_mod_cast h + | neg ih n => simpa only [preΨ_neg, neg_ne_zero] + using ih n <| neg_ne_zero.mp <| by exact_mod_cast h end preΨ @@ -327,14 +328,14 @@ private lemma natDegree_coeff_ΨSq_ofNat (n : ℕ) : have hd : (n + 1) ^ 2 - 1 = 2 * expDegree (n + 1) + if Even (n + 1) then 3 else 0 := by push_cast [← @Nat.cast_inj ℤ, add_sq, expDegree_cast (by omega : n + 1 ≠ 0)] split_ifs <;> ring1 - have hc : (n + 1) ^ 2 = expCoeff (n + 1) ^ 2 * if Even (n + 1) then 4 else 1 := by + have hc : (n + 1 : ℕ) ^ 2 = expCoeff (n + 1) ^ 2 * if Even (n + 1) then 4 else 1 := by push_cast [← @Int.cast_inj ℚ, expCoeff_cast] split_ifs <;> ring1 rw [ΨSq_ofNat, hd] constructor · refine natDegree_mul_le_of_le (dp h.1) ?_ split_ifs <;> simp only [apply_ite natDegree, natDegree_one.le, W.natDegree_Ψ₂Sq_le] - · erw [coeff_mul_of_natDegree_le (dp h.1), coeff_pow_of_natDegree_le h.1, h.2, apply_ite₂ coeff, + · rw [coeff_mul_of_natDegree_le (dp h.1), coeff_pow_of_natDegree_le h.1, h.2, apply_ite₂ coeff, coeff_Ψ₂Sq, coeff_one_zero, hc] · norm_cast split_ifs <;> simp only [apply_ite natDegree, natDegree_one.le, W.natDegree_Ψ₂Sq_le] @@ -342,13 +343,13 @@ private lemma natDegree_coeff_ΨSq_ofNat (n : ℕ) : lemma natDegree_ΨSq_le (n : ℤ) : (W.ΨSq n).natDegree ≤ n.natAbs ^ 2 - 1 := by induction n using Int.negInduction with | nat n => exact (W.natDegree_coeff_ΨSq_ofNat n).left - | neg => rwa [ΨSq_neg, Int.natAbs_neg] + | neg ih => simp only [ΨSq_neg, Int.natAbs_neg, ih] @[simp] lemma coeff_ΨSq (n : ℤ) : (W.ΨSq n).coeff (n.natAbs ^ 2 - 1) = n ^ 2 := by induction n using Int.negInduction with | nat n => exact_mod_cast (W.natDegree_coeff_ΨSq_ofNat n).right - | neg => rwa [ΨSq_neg, Int.natAbs_neg, ← Int.cast_pow, neg_sq, Int.cast_pow] + | neg ih => simp_rw [ΨSq_neg, Int.natAbs_neg, ← Int.cast_pow, neg_sq, Int.cast_pow, ih] lemma coeff_ΨSq_ne_zero [NoZeroDivisors R] {n : ℤ} (h : (n : R) ≠ 0) : (W.ΨSq n).coeff (n.natAbs ^ 2 - 1) ≠ 0 := by @@ -404,7 +405,7 @@ private lemma natDegree_coeff_Φ_ofNat (n : ℕ) : expCoeff (n + 3) * expCoeff (n + 1) * (if Even (n + 1) then 4 else 1) := by push_cast [← @Int.cast_inj ℚ, expCoeff_cast, Nat.even_add_one, ite_not] split_ifs <;> ring1 - erw [Φ_ofNat] + rw [Nat.cast_add, Nat.cast_one, Φ_ofNat] constructor · nth_rw 1 [← max_self <| (_ + _) ^ 2, hd, hd'] refine natDegree_sub_le_of_le (dm (dm natDegree_X_le (dp h.1)) ?_) (dm (dm h.1 h.1) ?_) @@ -419,13 +420,13 @@ private lemma natDegree_coeff_Φ_ofNat (n : ℕ) : lemma natDegree_Φ_le (n : ℤ) : (W.Φ n).natDegree ≤ n.natAbs ^ 2 := by induction n using Int.negInduction with | nat n => exact (W.natDegree_coeff_Φ_ofNat n).left - | neg => rwa [Φ_neg, Int.natAbs_neg] + | neg ih => simp only [Φ_neg, Int.natAbs_neg, ih] @[simp] lemma coeff_Φ (n : ℤ) : (W.Φ n).coeff (n.natAbs ^ 2) = 1 := by induction n using Int.negInduction with | nat n => exact (W.natDegree_coeff_Φ_ofNat n).right - | neg => rwa [Φ_neg, Int.natAbs_neg] + | neg ih => simp only [Φ_neg, Int.natAbs_neg, ih] lemma coeff_Φ_ne_zero [Nontrivial R] (n : ℤ) : (W.Φ n).coeff (n.natAbs ^ 2) ≠ 0 := W.coeff_Φ n ▸ one_ne_zero diff --git a/Mathlib/Data/Int/Defs.lean b/Mathlib/Data/Int/Defs.lean index 9fd2c581ce8e9..5fc0b54bd9a45 100644 --- a/Mathlib/Data/Int/Defs.lean +++ b/Mathlib/Data/Int/Defs.lean @@ -268,9 +268,9 @@ end inductionOn' /-- Inductively define a function on `ℤ` by defining it on `ℕ` and extending it from `n` to `-n`. -/ @[elab_as_elim] protected def negInduction {C : ℤ → Sort*} (nat : ∀ n : ℕ, C n) - (neg : ∀ n : ℕ, C n → C (-n)) : ∀ n : ℤ, C n + (neg : (∀ n : ℕ, C n) → ∀ n : ℕ, C (-n)) : ∀ n : ℤ, C n | .ofNat n => nat n - | .negSucc n => neg _ <| nat <| n + 1 + | .negSucc n => neg nat <| n + 1 /-- See `Int.inductionOn'` for an induction in both directions. -/ protected lemma le_induction {P : ℤ → Prop} {m : ℤ} (h0 : P m) diff --git a/Mathlib/Data/Nat/Cast/Defs.lean b/Mathlib/Data/Nat/Cast/Defs.lean index 6ef4845951ac7..b0606a4832ffc 100644 --- a/Mathlib/Data/Nat/Cast/Defs.lean +++ b/Mathlib/Data/Nat/Cast/Defs.lean @@ -166,6 +166,10 @@ theorem binCast_eq [AddMonoidWithOne R] (n : ℕ) : theorem cast_two [AddMonoidWithOne R] : ((2 : ℕ) : R) = (2 : R) := rfl +theorem cast_three [AddMonoidWithOne R] : ((3 : ℕ) : R) = (3 : R) := rfl + +theorem cast_four [AddMonoidWithOne R] : ((4 : ℕ) : R) = (4 : R) := rfl + attribute [simp, norm_cast] Int.natAbs_ofNat end Nat diff --git a/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean b/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean index 541332309d477..fb438770af84a 100644 --- a/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean +++ b/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean @@ -14,10 +14,10 @@ This file defines the type of an elliptic divisibility sequence (EDS) and a few ## Mathematical background -Let $R$ be a commutative ring. An elliptic sequence is a sequence $W : \mathbb{Z} \to R$ satisfying -$$ W(m + n)W(m - n)W(r)^2 = W(m + r)W(m - r)W(n)^2 - W(n + r)W(n - r)W(m)^2, $$ -for any $m, n, r \in \mathbb{Z}$. A divisibility sequence is a sequence $W : \mathbb{Z} \to R$ -satisfying $W(m) \mid W(n)$ for any $m, n \in \mathbb{Z}$ such that $m \mid n$. +Let `R` be a commutative ring. An elliptic sequence is a sequence `W : ℤ → R` satisfying +`W(m + n)W(m - n)W(r)² = W(m + r)W(m - r)W(n)² - W(n + r)W(n - r)W(m)²` for any `m, n, r ∈ ℤ`. +A divisibility sequence is a sequence `W : ℤ → R` satisfying `W(m) ∣ W(n)` for any `m, n ∈ ℤ` such +that `m ∣ n`. An elliptic divisibility sequence is simply a divisibility sequence that is elliptic. Some examples of EDSs include * the identity sequence, @@ -50,8 +50,8 @@ One reason is to avoid the necessity for ring division by `b` in the inductive d `normEDS b c d (2 * (m + 3))`. The idea is that, it can be shown that `normEDS b c d (2 * (m + 3))` always contains a factor of `b`, so it is possible to remove a factor of `b` *a posteriori*, but stating this lemma requires first defining `normEDS b c d (2 * (m + 3))`, which requires having this -factor of `b` *a priori*. Another reason is to allow the definition of univariate $n$-division -polynomials of elliptic curves, omitting a factor of the bivariate $2$-division polynomial. +factor of `b` *a priori*. Another reason is to allow the definition of univariate `n`-division +polynomials of elliptic curves, omitting a factor of the bivariate `2`-division polynomial. ## References @@ -62,9 +62,13 @@ M Ward, *Memoir on Elliptic Divisibility Sequences* elliptic, divisibility, sequence -/ -universe u v w +universe u v -variable {R : Type u} {S : Type v} [CommRing R] [CommRing S] (W : ℤ → R) (f : R →+* S) +variable {R : Type u} [CommRing R] + +section IsEllDivSequence + +variable (W : ℤ → R) /-- The proposition that a sequence indexed by integers is an elliptic sequence. -/ def IsEllSequence : Prop := @@ -101,6 +105,42 @@ lemma IsDivSequence.smul (h : IsDivSequence W) (x : R) : IsDivSequence (x • W) lemma IsEllDivSequence.smul (h : IsEllDivSequence W) (x : R) : IsEllDivSequence (x • W) := ⟨h.left.smul x, h.right.smul x⟩ +end IsEllDivSequence + +/-- Strong recursion principle for a normalised EDS: if we have + * `P 0`, `P 1`, `P 2`, `P 3`, and `P 4`, + * for all `m : ℕ` we can prove `P (2 * (m + 3))` from `P k` for all `k < 2 * (m + 3)`, and + * for all `m : ℕ` we can prove `P (2 * (m + 2) + 1)` from `P k` for all `k < 2 * (m + 2) + 1`, +then we have `P n` for all `n : ℕ`. -/ +@[elab_as_elim] +noncomputable def normEDSRec' {P : ℕ → Sort u} + (zero : P 0) (one : P 1) (two : P 2) (three : P 3) (four : P 4) + (even : ∀ m : ℕ, (∀ k < 2 * (m + 3), P k) → P (2 * (m + 3))) + (odd : ∀ m : ℕ, (∀ k < 2 * (m + 2) + 1, P k) → P (2 * (m + 2) + 1)) (n : ℕ) : P n := + n.evenOddStrongRec (by rintro (_ | _ | _ | _) h; exacts [zero, two, four, even _ h]) + (by rintro (_ | _ | _) h; exacts [one, three, odd _ h]) + +/-- Recursion principle for a normalised EDS: if we have + * `P 0`, `P 1`, `P 2`, `P 3`, and `P 4`, + * for all `m : ℕ` we can prove `P (2 * (m + 3))` from `P (m + 1)`, `P (m + 2)`, `P (m + 3)`, + `P (m + 4)`, and `P (m + 5)`, and + * for all `m : ℕ` we can prove `P (2 * (m + 2) + 1)` from `P (m + 1)`, `P (m + 2)`, `P (m + 3)`, + and `P (m + 4)`, +then we have `P n` for all `n : ℕ`. -/ +@[elab_as_elim] +noncomputable def normEDSRec {P : ℕ → Sort u} + (zero : P 0) (one : P 1) (two : P 2) (three : P 3) (four : P 4) + (even : ∀ m : ℕ, P (m + 1) → P (m + 2) → P (m + 3) → P (m + 4) → P (m + 5) → P (2 * (m + 3))) + (odd : ∀ m : ℕ, P (m + 1) → P (m + 2) → P (m + 3) → P (m + 4) → P (2 * (m + 2) + 1)) (n : ℕ) : + P n := + normEDSRec' zero one two three four + (fun _ ih => by apply even <;> exact ih _ <| by linarith only) + (fun _ ih => by apply odd <;> exact ih _ <| by linarith only) n + +variable (b c d : R) + +section PreNormEDS + /-- The auxiliary sequence for a normalised EDS `W : ℕ → R`, with initial values `W(0) = 0`, `W(1) = 1`, `W(2) = 1`, `W(3) = c`, and `W(4) = d` and extra parameter `b`. -/ def preNormEDS' (b c d : R) : ℕ → R @@ -123,8 +163,6 @@ def preNormEDS' (b c d : R) : ℕ → R preNormEDS' b c d (m + 2) ^ 2 * preNormEDS' b c d (m + 3) * preNormEDS' b c d (m + 5) - preNormEDS' b c d (m + 1) * preNormEDS' b c d (m + 3) * preNormEDS' b c d (m + 4) ^ 2 -variable (b c d : R) - @[simp] lemma preNormEDS'_zero : preNormEDS' b c d 0 = 0 := by rw [preNormEDS'] @@ -191,22 +229,64 @@ lemma preNormEDS_three : preNormEDS b c d 3 = c := by lemma preNormEDS_four : preNormEDS b c d 4 = d := by erw [preNormEDS_ofNat, preNormEDS'_four] -lemma preNormEDS_odd (m : ℕ) : preNormEDS b c d (2 * (m + 2) + 1) = - preNormEDS b c d (m + 4) * preNormEDS b c d (m + 2) ^ 3 * (if Even m then b else 1) - - preNormEDS b c d (m + 1) * preNormEDS b c d (m + 3) ^ 3 * (if Even m then 1 else b) := by - repeat erw [preNormEDS_ofNat] - exact preNormEDS'_odd .. - -lemma preNormEDS_even (m : ℕ) : preNormEDS b c d (2 * (m + 3)) = +lemma preNormEDS_even_ofNat (m : ℕ) : preNormEDS b c d (2 * (m + 3)) = preNormEDS b c d (m + 2) ^ 2 * preNormEDS b c d (m + 3) * preNormEDS b c d (m + 5) - preNormEDS b c d (m + 1) * preNormEDS b c d (m + 3) * preNormEDS b c d (m + 4) ^ 2 := by repeat erw [preNormEDS_ofNat] exact preNormEDS'_even .. +lemma preNormEDS_odd_ofNat (m : ℕ) : preNormEDS b c d (2 * (m + 2) + 1) = + preNormEDS b c d (m + 4) * preNormEDS b c d (m + 2) ^ 3 * (if Even m then b else 1) - + preNormEDS b c d (m + 1) * preNormEDS b c d (m + 3) ^ 3 * (if Even m then 1 else b) := by + repeat erw [preNormEDS_ofNat] + exact preNormEDS'_odd .. + @[simp] lemma preNormEDS_neg (n : ℤ) : preNormEDS b c d (-n) = -preNormEDS b c d n := by rw [preNormEDS, Int.sign_neg, Int.cast_neg, neg_mul, Int.natAbs_neg, preNormEDS] +lemma preNormEDS_even (m : ℤ) : preNormEDS b c d (2 * m) = + preNormEDS b c d (m - 1) ^ 2 * preNormEDS b c d m * preNormEDS b c d (m + 2) - + preNormEDS b c d (m - 2) * preNormEDS b c d m * preNormEDS b c d (m + 1) ^ 2 := by + induction m using Int.negInduction with + | nat m => + rcases m with _ | _ | _ | m + · simp + · simp + · simp + · simp only [Int.natCast_add, Nat.cast_one] + rw [Int.add_sub_cancel, show (m : ℤ) + 1 + 1 + 1 = m + 1 + 2 by rfl, Int.add_sub_cancel] + exact preNormEDS_even_ofNat .. + | neg h m => + simp_rw [show 2 * -(m : ℤ) = -(2 * m) by omega, show -(m : ℤ) - 1 = -(m + 1) by omega, + show -(m : ℤ) + 2 = -(m - 2) by omega, show -(m : ℤ) - 2 = -(m + 2) by omega, + show -(m : ℤ) + 1 = -(m - 1) by omega, preNormEDS_neg, h m] + ring1 + +lemma preNormEDS_odd (m : ℤ) : preNormEDS b c d (2 * m + 1) = + preNormEDS b c d (m + 2) * preNormEDS b c d m ^ 3 * (if Even m then b else 1) - + preNormEDS b c d (m - 1) * preNormEDS b c d (m + 1) ^ 3 * (if Even m then 1 else b) := by + induction m using Int.negInduction with + | nat m => + rcases m with _ | _ | m + · simp + · simp + · simp only [Int.natCast_add, Nat.cast_one, Int.even_add_one, not_not, Int.even_coe_nat] + rw [Int.add_sub_cancel] + exact preNormEDS_odd_ofNat .. + | neg h m => + rcases m with _ | m + · simp + · simp_rw [Int.natCast_add, Nat.cast_one, show 2 * -(m + 1 : ℤ) + 1 = -(2 * m + 1) by rfl, + show -(m + 1 : ℤ) + 2 = -(m - 1) by omega, show -(m + 1 : ℤ) - 1 = -(m + 2) by rfl, + show -(m + 1 : ℤ) + 1 = -m by omega, preNormEDS_neg, even_neg, Int.even_add_one, ite_not, + h m] + ring1 + +end PreNormEDS + +section NormEDS + /-- The canonical example of a normalised EDS `W : ℤ → R`, with initial values `W(0) = 0`, `W(1) = 1`, `W(2) = b`, `W(3) = c`, and `W(4) = d * b`. @@ -239,53 +319,48 @@ lemma normEDS_three : normEDS b c d 3 = c := by lemma normEDS_four : normEDS b c d 4 = d * b := by erw [normEDS_ofNat, preNormEDS'_four, if_pos <| by decide] -lemma normEDS_odd (m : ℕ) : normEDS b c d (2 * (m + 2) + 1) = - normEDS b c d (m + 4) * normEDS b c d (m + 2) ^ 3 - - normEDS b c d (m + 1) * normEDS b c d (m + 3) ^ 3 := by - repeat erw [normEDS_ofNat] - simp_rw [preNormEDS'_odd, if_neg (m + 2).not_even_two_mul_add_one, Nat.even_add_one, ite_not] - split_ifs <;> ring1 - -lemma normEDS_even (m : ℕ) : normEDS b c d (2 * (m + 3)) * b = +lemma normEDS_even_ofNat (m : ℕ) : normEDS b c d (2 * (m + 3)) * b = normEDS b c d (m + 2) ^ 2 * normEDS b c d (m + 3) * normEDS b c d (m + 5) - normEDS b c d (m + 1) * normEDS b c d (m + 3) * normEDS b c d (m + 4) ^ 2 := by repeat erw [normEDS_ofNat] simp only [preNormEDS'_even, if_pos <| even_two_mul _, Nat.even_add_one, ite_not] split_ifs <;> ring1 +lemma normEDS_odd_ofNat (m : ℕ) : normEDS b c d (2 * (m + 2) + 1) = + normEDS b c d (m + 4) * normEDS b c d (m + 2) ^ 3 - + normEDS b c d (m + 1) * normEDS b c d (m + 3) ^ 3 := by + repeat erw [normEDS_ofNat] + simp_rw [preNormEDS'_odd, if_neg (m + 2).not_even_two_mul_add_one, Nat.even_add_one, ite_not] + split_ifs <;> ring1 + @[simp] lemma normEDS_neg (n : ℤ) : normEDS b c d (-n) = -normEDS b c d n := by simp only [normEDS, preNormEDS_neg, neg_mul, even_neg] -/-- Strong recursion principle for a normalised EDS: if we have - * `P 0`, `P 1`, `P 2`, `P 3`, and `P 4`, - * for all `m : ℕ` we can prove `P (2 * (m + 3))` from `P k` for all `k < 2 * (m + 3)`, and - * for all `m : ℕ` we can prove `P (2 * (m + 2) + 1)` from `P k` for all `k < 2 * (m + 2) + 1`, -then we have `P n` for all `n : ℕ`. -/ -@[elab_as_elim] -noncomputable def normEDSRec' {P : ℕ → Sort u} - (zero : P 0) (one : P 1) (two : P 2) (three : P 3) (four : P 4) - (even : ∀ m : ℕ, (∀ k < 2 * (m + 3), P k) → P (2 * (m + 3))) - (odd : ∀ m : ℕ, (∀ k < 2 * (m + 2) + 1, P k) → P (2 * (m + 2) + 1)) (n : ℕ) : P n := - n.evenOddStrongRec (by rintro (_ | _ | _ | _) h; exacts [zero, two, four, even _ h]) - (by rintro (_ | _ | _) h; exacts [one, three, odd _ h]) +lemma normEDS_even (m : ℤ) : normEDS b c d (2 * m) * b = + normEDS b c d (m - 1) ^ 2 * normEDS b c d m * normEDS b c d (m + 2) - + normEDS b c d (m - 2) * normEDS b c d m * normEDS b c d (m + 1) ^ 2 := by + simp only [normEDS, preNormEDS_even, if_pos <| even_two_mul m, show m + 2 = m + 1 + 1 by ring1, + Int.even_add_one, show m - 2 = m - 1 - 1 by ring1, Int.even_sub_one, ite_not] + by_cases hm : Even m + · simp only [if_pos hm] + ring1 + · simp only [if_neg hm] + ring1 -/-- Recursion principle for a normalised EDS: if we have - * `P 0`, `P 1`, `P 2`, `P 3`, and `P 4`, - * for all `m : ℕ` we can prove `P (2 * (m + 3))` from `P (m + 1)`, `P (m + 2)`, `P (m + 3)`, - `P (m + 4)`, and `P (m + 5)`, and - * for all `m : ℕ` we can prove `P (2 * (m + 2) + 1)` from `P (m + 1)`, `P (m + 2)`, `P (m + 3)`, - and `P (m + 4)`, -then we have `P n` for all `n : ℕ`. -/ -@[elab_as_elim] -noncomputable def normEDSRec {P : ℕ → Sort u} - (zero : P 0) (one : P 1) (two : P 2) (three : P 3) (four : P 4) - (even : ∀ m : ℕ, P (m + 1) → P (m + 2) → P (m + 3) → P (m + 4) → P (m + 5) → P (2 * (m + 3))) - (odd : ∀ m : ℕ, P (m + 1) → P (m + 2) → P (m + 3) → P (m + 4) → P (2 * (m + 2) + 1)) (n : ℕ) : - P n := - normEDSRec' zero one two three four - (fun _ ih => by apply even <;> exact ih _ <| by linarith only) - (fun _ ih => by apply odd <;> exact ih _ <| by linarith only) n +lemma normEDS_odd (m : ℤ) : normEDS b c d (2 * m + 1) = + normEDS b c d (m + 2) * normEDS b c d m ^ 3 - + normEDS b c d (m - 1) * normEDS b c d (m + 1) ^ 3 := by + simp only [normEDS, preNormEDS_odd, if_neg m.not_even_two_mul_add_one] + conv_lhs => rw [← @one_pow R _ 4, ← ite_pow, ← ite_pow] + simp only [show m + 2 = m + 1 + 1 by ring1, Int.even_add_one, Int.even_sub_one, ite_not] + ring1 + +end NormEDS + +section Map + +variable {S : Type v} [CommRing S] (f : R →+* S) lemma map_preNormEDS' (n : ℕ) : f (preNormEDS' b c d n) = preNormEDS' (f b) (f c) (f d) n := by induction n using normEDSRec' with @@ -303,3 +378,5 @@ lemma map_preNormEDS (n : ℤ) : f (preNormEDS b c d n) = preNormEDS (f b) (f c) lemma map_normEDS (n : ℤ) : f (normEDS b c d n) = normEDS (f b) (f c) (f d) n := by rw [normEDS, map_mul, map_preNormEDS, map_pow, apply_ite f, map_one, normEDS] + +end Map From 496527772ac807ef3d1ddd4868f2048f9ea478e1 Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Fri, 29 Nov 2024 12:22:53 +0000 Subject: [PATCH 416/829] feat(AlgebraicGeometry): the small site associated to a morphism property (#18945) Given a sufficiently good morphism property `P`, we define the induced topologies on the category of `S`-schemes and on the category of `S`-schemes with `P`. --- Mathlib.lean | 1 + Mathlib/AlgebraicGeometry/Sites/Small.lean | 260 +++++++++++++++++++++ Mathlib/CategoryTheory/Sites/Over.lean | 7 + Mathlib/CategoryTheory/Sites/Sieves.lean | 17 ++ 4 files changed, 285 insertions(+) create mode 100644 Mathlib/AlgebraicGeometry/Sites/Small.lean diff --git a/Mathlib.lean b/Mathlib.lean index a6e93d1a935f9..d8c2ef98c4eed 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1006,6 +1006,7 @@ import Mathlib.AlgebraicGeometry.Scheme import Mathlib.AlgebraicGeometry.Sites.BigZariski import Mathlib.AlgebraicGeometry.Sites.Etale import Mathlib.AlgebraicGeometry.Sites.MorphismProperty +import Mathlib.AlgebraicGeometry.Sites.Small import Mathlib.AlgebraicGeometry.Spec import Mathlib.AlgebraicGeometry.SpreadingOut import Mathlib.AlgebraicGeometry.Stalk diff --git a/Mathlib/AlgebraicGeometry/Sites/Small.lean b/Mathlib/AlgebraicGeometry/Sites/Small.lean new file mode 100644 index 0000000000000..5e5bfcaf299da --- /dev/null +++ b/Mathlib/AlgebraicGeometry/Sites/Small.lean @@ -0,0 +1,260 @@ +/- +Copyright (c) 2024 Christian Merten. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Christian Merten +-/ +import Mathlib.AlgebraicGeometry.Cover.Over +import Mathlib.AlgebraicGeometry.Sites.MorphismProperty +import Mathlib.CategoryTheory.Sites.DenseSubsite.InducedTopology +import Mathlib.CategoryTheory.Sites.Over + +/-! +# Small sites + +In this file we define the small sites associated to morphism properties and give +generating pretopologies. + +## Main definitions + +- `AlgebraicGeometry.Scheme.overGrothendieckTopology`: the Grothendieck topology on `Over S` + obtained by localizing the topology on `Scheme` induced by `P` at `S`. +- `AlgebraicGeometry.Scheme.overPretopology`: the pretopology on `Over S` defined by + `P`-coverings of `S`-schemes. The induced topology agrees with + `AlgebraicGeometry.Scheme.overGrothendieckTopology`. +- `AlgebraicGeometry.Scheme.smallGrothendieckTopology`: the by the inclusion + `P.Over ⊤ S ⥤ Over S` induced topology on `P.Over ⊤ S`. +- `AlgebraicGeometry.Scheme.smallPretopology`: the pretopology on `P.Over ⊤ S` defined by + `P`-coverings of `S`-schemes with `P`. The induced topology agrees + with `AlgebraicGeometry.Scheme.smallGrothendieckTopology`. + +-/ + +universe v u + +open CategoryTheory Limits + +namespace AlgebraicGeometry.Scheme + +variable {P Q : MorphismProperty Scheme.{u}} {S : Scheme.{u}} + +/-- The presieve defined by a `P`-cover of `S`-schemes. -/ +def Cover.toPresieveOver {X : Over S} (𝒰 : Cover.{u} P X.left) [𝒰.Over S] : Presieve X := + Presieve.ofArrows (fun i ↦ (𝒰.obj i).asOver S) (fun i ↦ (𝒰.map i).asOver S) + +/-- The presieve defined by a `P`-cover of `S`-schemes with `Q`. -/ +def Cover.toPresieveOverProp {X : Q.Over ⊤ S} (𝒰 : Cover.{u} P X.left) [𝒰.Over S] + (h : ∀ j, Q (𝒰.obj j ↘ S)) : Presieve X := + Presieve.ofArrows (fun i ↦ (𝒰.obj i).asOverProp S (h i)) (fun i ↦ (𝒰.map i).asOverProp S) + +lemma Cover.overEquiv_generate_toPresieveOver_eq_ofArrows {X : Over S} (𝒰 : Cover.{u} P X.left) + [𝒰.Over S] : Sieve.overEquiv X (Sieve.generate 𝒰.toPresieveOver) = + Sieve.ofArrows 𝒰.obj 𝒰.map := by + ext V f + simp only [Sieve.overEquiv_iff, Functor.const_obj_obj, Sieve.generate_apply] + constructor + · rintro ⟨U, h, g, ⟨k⟩, hcomp⟩ + exact ⟨𝒰.obj k, h.left, 𝒰.map k, ⟨k⟩, congrArg CommaMorphism.left hcomp⟩ + · rintro ⟨U, h, g, ⟨k⟩, hcomp⟩ + have : 𝒰.map k ≫ X.hom = 𝒰.obj k ↘ S := comp_over (𝒰.map k) S + refine ⟨(𝒰.obj k).asOver S, Over.homMk h (by simp [← hcomp, this]), (𝒰.map k).asOver S, ⟨k⟩, ?_⟩ + ext : 1 + simpa + +lemma Cover.toPresieveOver_le_arrows_iff {X : Over S} (R : Sieve X) (𝒰 : Cover.{u} P X.left) + [𝒰.Over S] : + 𝒰.toPresieveOver ≤ R.arrows ↔ + Presieve.ofArrows 𝒰.obj 𝒰.map ≤ (Sieve.overEquiv X R).arrows := by + simp_rw [← Sieve.giGenerate.gc.le_iff_le, ← Sieve.overEquiv_le_overEquiv_iff] + rw [overEquiv_generate_toPresieveOver_eq_ofArrows] + +variable [P.IsMultiplicative] [P.RespectsIso] + [P.IsStableUnderBaseChange] [IsJointlySurjectivePreserving P] + +variable (P Q S) + +/-- The pretopology on `Over S` induced by `P` where coverings are given by `P`-covers +of `S`-schemes. -/ +def overPretopology : Pretopology (Over S) where + coverings Y R := ∃ (𝒰 : Cover.{u} P Y.left) (_ : 𝒰.Over S), R = 𝒰.toPresieveOver + has_isos {X Y} f _ := ⟨coverOfIsIso f.left, inferInstance, (Presieve.ofArrows_pUnit _).symm⟩ + pullbacks := by + rintro Y X f _ ⟨𝒰, h, rfl⟩ + refine ⟨𝒰.pullbackCoverOver' S f.left, inferInstance, ?_⟩ + simpa [Cover.toPresieveOver] using + (Presieve.ofArrows_pullback f (fun i ↦ (𝒰.obj i).asOver S) (fun i ↦ (𝒰.map i).asOver S)).symm + transitive := by + rintro X _ T ⟨𝒰, h, rfl⟩ H + choose V h hV using H + refine ⟨𝒰.bind (fun j => V ((𝒰.map j).asOver S) ⟨j⟩), inferInstance, ?_⟩ + convert Presieve.ofArrows_bind _ (fun j ↦ (𝒰.map j).asOver S) _ + (fun Y f H j ↦ ((V f H).obj j).asOver S) (fun Y f H j ↦ ((V f H).map j).asOver S) + apply hV + +/-- The topology on `Over S` induced from the topology on `Scheme` defined by `P`. +This agrees with the topology induced by `S.overPretopology P`, see +`AlgebraicGeometry.Scheme.overGrothendieckTopology_eq_toGrothendieck_overPretopology`. -/ +abbrev overGrothendieckTopology : GrothendieckTopology (Over S) := + (Scheme.grothendieckTopology P).over S + +lemma overGrothendieckTopology_eq_toGrothendieck_overPretopology : + S.overGrothendieckTopology P = (S.overPretopology P).toGrothendieck := by + ext X R + rw [GrothendieckTopology.mem_over_iff, Pretopology.mem_toGrothendieck] + constructor + · rintro ⟨T, ⟨𝒰, rfl⟩, hT⟩ + letI (i : 𝒰.J) : (𝒰.obj i).Over S := { hom := 𝒰.map i ≫ X.hom } + letI : 𝒰.Over S := + { over := inferInstance + isOver_map := fun i ↦ ⟨rfl⟩ } + use 𝒰.toPresieveOver, ⟨𝒰, inferInstance, rfl⟩ + rwa [Cover.toPresieveOver_le_arrows_iff] + · rintro ⟨T, ⟨𝒰, h, rfl⟩, hT⟩ + use Presieve.ofArrows 𝒰.obj 𝒰.map, ⟨𝒰, rfl⟩ + rwa [Cover.toPresieveOver_le_arrows_iff] at hT + +variable {S} + +lemma mem_overGrothendieckTopology (X : Over S) (R : Sieve X) : + R ∈ S.overGrothendieckTopology P X ↔ + ∃ (𝒰 : Cover.{u} P X.left) (_ : 𝒰.Over S), 𝒰.toPresieveOver ≤ R.arrows := by + rw [overGrothendieckTopology_eq_toGrothendieck_overPretopology] + constructor + · rintro ⟨T, ⟨𝒰, h, rfl⟩, hle⟩ + use 𝒰, h + · rintro ⟨𝒰, h𝒰, hle⟩ + exact ⟨𝒰.toPresieveOver, ⟨𝒰, h𝒰, rfl⟩, hle⟩ + +variable [Q.IsStableUnderComposition] + +variable (S) {P Q} in +lemma locallyCoverDense_of_le (hPQ : P ≤ Q) : + (MorphismProperty.Over.forget Q ⊤ S).LocallyCoverDense (overGrothendieckTopology P S) where + functorPushforward_functorPullback_mem X := by + intro ⟨T, hT⟩ + rw [mem_overGrothendieckTopology] at hT ⊢ + obtain ⟨𝒰, h, hle⟩ := hT + use 𝒰, h + rintro - - ⟨i⟩ + have p : Q (𝒰.obj i ↘ S) := by + rw [← comp_over (𝒰.map i) S] + exact Q.comp_mem _ _ (hPQ _ <| 𝒰.map_prop i) X.prop + use (𝒰.obj i).asOverProp S p, MorphismProperty.Over.homMk (𝒰.map i) (comp_over (𝒰.map i) S), 𝟙 _ + exact ⟨hle _ ⟨i⟩, rfl⟩ + +instance : (MorphismProperty.Over.forget P ⊤ S).LocallyCoverDense (overGrothendieckTopology P S) := + locallyCoverDense_of_le S le_rfl + +variable (S) {P Q} in +/-- If `P` and `Q` are morphism properties with `P ≤ Q`, this is the Grothendieck topology +induced via the forgetful functor `Q.Over ⊤ S ⥤ Over S` by the topology defined by `P`. -/ +abbrev smallGrothendieckTopologyOfLE (hPQ : P ≤ Q) : GrothendieckTopology (Q.Over ⊤ S) := + letI : (MorphismProperty.Over.forget Q ⊤ S).LocallyCoverDense (overGrothendieckTopology P S) := + locallyCoverDense_of_le S hPQ + (MorphismProperty.Over.forget Q ⊤ S).inducedTopology (S.overGrothendieckTopology P) + +/-- The Grothendieck topology on the category of schemes over `S` with `P` induced by `P`, i.e. +coverings are simply surjective families. This is the induced topology by the topology on `S` +defined by `P` via the inclusion `P.Over ⊤ S ⥤ Over S`. + +This is a special case of `smallGrothendieckTopologyOfLE` for the case `P = Q`. -/ +abbrev smallGrothendieckTopology : GrothendieckTopology (P.Over ⊤ S) := + (MorphismProperty.Over.forget P ⊤ S).inducedTopology (S.overGrothendieckTopology P) + +variable [Q.IsStableUnderBaseChange] [Q.HasOfPostcompProperty Q] + +/-- The pretopology defind on the subcategory of `S`-schemes satisfying `Q` where coverings +are given by `P`-coverings in `S`-schemes satisfying `Q`. +The most common case is `P = Q`. In this case, this is simply surjective families +in `S`-schemes with `P`. -/ +def smallPretopology : Pretopology (Q.Over ⊤ S) where + coverings Y R := ∃ (𝒰 : Cover.{u} P Y.left) (_ : 𝒰.Over S) (h : ∀ j : 𝒰.J, Q (𝒰.obj j ↘ S)), + R = 𝒰.toPresieveOverProp h + has_isos {X Y} f := ⟨coverOfIsIso f.left, inferInstance, fun _ ↦ Y.prop, + (Presieve.ofArrows_pUnit _).symm⟩ + pullbacks := by + rintro Y X f _ ⟨𝒰, h, p, rfl⟩ + refine ⟨𝒰.pullbackCoverOverProp' S f.left (Q := Q) Y.prop X.prop p, inferInstance, ?_, ?_⟩ + · intro j + apply MorphismProperty.Comma.prop + · exact (Presieve.ofArrows_pullback f (fun i ↦ ⟨(𝒰.obj i).asOver S, p i⟩) + (fun i ↦ ⟨(𝒰.map i).asOver S, trivial, trivial⟩)).symm + transitive := by + rintro X _ T ⟨𝒰, h, p, rfl⟩ H + choose V h pV hV using H + let 𝒱j (j : 𝒰.J) : (Cover P ((𝒰.obj j).asOverProp S (p j)).left) := + V ((𝒰.map j).asOverProp S) ⟨j⟩ + refine ⟨𝒰.bind (fun j ↦ 𝒱j j), inferInstance, fun j ↦ pV _ _ _, ?_⟩ + convert Presieve.ofArrows_bind _ (fun j ↦ ((𝒰.map j).asOverProp S)) _ + (fun Y f H j ↦ ((V f H).obj j).asOverProp S (pV _ _ _)) + (fun Y f H j ↦ ((V f H).map j).asOverProp S) + apply hV + +variable (S) {P Q} in +lemma smallGrothendieckTopologyOfLE_eq_toGrothendieck_smallPretopology (hPQ : P ≤ Q) : + S.smallGrothendieckTopologyOfLE hPQ = (S.smallPretopology P Q).toGrothendieck := by + ext X R + simp only [Pretopology.mem_toGrothendieck, Functor.mem_inducedTopology_sieves_iff, + MorphismProperty.Comma.forget_obj, mem_overGrothendieckTopology] + constructor + · intro ⟨𝒰, h, le⟩ + have hj (j : 𝒰.J) : Q (𝒰.obj j ↘ S) := by + rw [← comp_over (𝒰.map j)] + exact Q.comp_mem _ _ (hPQ _ <| 𝒰.map_prop _) X.prop + refine ⟨𝒰.toPresieveOverProp hj, ?_, ?_⟩ + · use 𝒰, h, hj + · rintro - - ⟨i⟩ + let fi : (𝒰.obj i).asOverProp S (hj i) ⟶ X := (𝒰.map i).asOverProp S + have : R.functorPushforward _ ((MorphismProperty.Over.forget Q ⊤ S).map fi) := le _ ⟨i⟩ + rwa [Sieve.functorPushforward_apply, + Sieve.mem_functorPushforward_iff_of_full_of_faithful] at this + · rintro ⟨T, ⟨𝒰, h, p, rfl⟩, le⟩ + use 𝒰, h + rintro - - ⟨i⟩ + exact ⟨(𝒰.obj i).asOverProp S (p i), (𝒰.map i).asOverProp S, 𝟙 _, le _ ⟨i⟩, rfl⟩ + +lemma smallGrothendieckTopology_eq_toGrothendieck_smallPretopology [P.HasOfPostcompProperty P] : + S.smallGrothendieckTopology P = (S.smallPretopology P P).toGrothendieck := + S.smallGrothendieckTopologyOfLE_eq_toGrothendieck_smallPretopology le_rfl + +variable {P Q} + +lemma mem_toGrothendieck_smallPretopology (X : Q.Over ⊤ S) (R : Sieve X) : + R ∈ (S.smallPretopology P Q).toGrothendieck _ X ↔ + ∀ x : X.left, ∃ (Y : Q.Over ⊤ S) (f : Y ⟶ X) (y : Y.left), + R f ∧ P f.left ∧ f.left.base y = x := by + rw [Pretopology.mem_toGrothendieck] + refine ⟨?_, fun h ↦ ?_⟩ + · rintro ⟨T, ⟨𝒰, h, p, rfl⟩, hle⟩ + intro x + obtain ⟨y, hy⟩ := 𝒰.covers x + refine ⟨(𝒰.obj (𝒰.f x)).asOverProp S (p _), (𝒰.map (𝒰.f x)).asOverProp S, y, hle _ ?_, + 𝒰.map_prop _, hy⟩ + use 𝒰.f x + · choose Y f y hf hP hy using h + let 𝒰 : X.left.Cover P := + { J := X.left, + obj := fun i ↦ (Y i).left + map := fun i ↦ (f i).left + map_prop := hP + f := id + covers := fun i ↦ ⟨y i, hy i⟩ } + letI : 𝒰.Over S := + { over := fun i ↦ inferInstance + isOver_map := fun i ↦ inferInstance } + refine ⟨𝒰.toPresieveOverProp fun i ↦ MorphismProperty.Comma.prop _, ?_, ?_⟩ + · use 𝒰, inferInstance, fun i ↦ MorphismProperty.Comma.prop _ + · rintro - - ⟨i⟩ + exact hf i + +lemma mem_smallGrothendieckTopology [P.HasOfPostcompProperty P] (X : P.Over ⊤ S) (R : Sieve X) : + R ∈ S.smallGrothendieckTopology P X ↔ + ∃ (𝒰 : Cover.{u} P X.left) (_ : 𝒰.Over S) (h : ∀ j, P (𝒰.obj j ↘ S)), + 𝒰.toPresieveOverProp h ≤ R.arrows := by + rw [smallGrothendieckTopology_eq_toGrothendieck_smallPretopology] + constructor + · rintro ⟨T, ⟨𝒰, h, p, rfl⟩, hle⟩ + use 𝒰, h, p + · rintro ⟨𝒰, h𝒰, p, hle⟩ + exact ⟨𝒰.toPresieveOverProp p, ⟨𝒰, h𝒰, p, rfl⟩, hle⟩ + +end AlgebraicGeometry.Scheme diff --git a/Mathlib/CategoryTheory/Sites/Over.lean b/Mathlib/CategoryTheory/Sites/Over.lean index f012575ed625c..2b06a21100f61 100644 --- a/Mathlib/CategoryTheory/Sites/Over.lean +++ b/Mathlib/CategoryTheory/Sites/Over.lean @@ -66,6 +66,13 @@ lemma overEquiv_symm_top {X : C} (Y : Over X) : (overEquiv Y).symm ⊤ = ⊤ := (overEquiv Y).injective (by simp) +lemma overEquiv_le_overEquiv_iff {X : C} {Y : Over X} (R₁ R₂ : Sieve Y) : + R₁.overEquiv Y ≤ R₂.overEquiv Y ↔ R₁ ≤ R₂ := by + refine ⟨fun h ↦ ?_, fun h ↦ Sieve.functorPushforward_monotone _ _ h⟩ + replace h : (overEquiv Y).symm (R₁.overEquiv Y) ≤ (overEquiv Y).symm (R₂.overEquiv Y) := + Sieve.functorPullback_monotone _ _ h + simpa using h + lemma overEquiv_pullback {X : C} {Y₁ Y₂ : Over X} (f : Y₁ ⟶ Y₂) (S : Sieve Y₂) : overEquiv _ (S.pullback f) = (overEquiv _ S).pullback f.left := by ext Z g diff --git a/Mathlib/CategoryTheory/Sites/Sieves.lean b/Mathlib/CategoryTheory/Sites/Sieves.lean index 66591c5a8f117..38b3981948c40 100644 --- a/Mathlib/CategoryTheory/Sites/Sieves.lean +++ b/Mathlib/CategoryTheory/Sites/Sieves.lean @@ -756,6 +756,23 @@ lemma pullback_functorPushforward_equivalence_eq {X : C} (S : Sieve X) : Sieve.pullback (e.unit.app X) (Sieve.functorPushforward e.inverse (Sieve.functorPushforward e.functor S)) = S := by ext; simp +lemma mem_functorPushforward_iff_of_full [F.Full] {X Y : C} (R : Sieve X) (f : F.obj Y ⟶ F.obj X) : + (R.arrows.functorPushforward F) f ↔ ∃ (g : Y ⟶ X), F.map g = f ∧ R g := by + refine ⟨fun ⟨Z, g, h, hg, hcomp⟩ ↦ ?_, fun ⟨g, hcomp, hg⟩ ↦ ?_⟩ + · obtain ⟨h', hh'⟩ := F.map_surjective h + use h' ≫ g + simp only [Functor.map_comp, hh', hcomp, true_and] + apply R.downward_closed hg + · use Y, g, 𝟙 _, hg + simp [hcomp] + +lemma mem_functorPushforward_iff_of_full_of_faithful [F.Full] [F.Faithful] + {X Y : C} (R : Sieve X) (f : Y ⟶ X) : + (R.arrows.functorPushforward F) (F.map f) ↔ R f := by + rw [Sieve.mem_functorPushforward_iff_of_full] + refine ⟨fun ⟨g, hcomp, hg⟩ ↦ ?_, fun hf ↦ ⟨f, rfl, hf⟩⟩ + rwa [← F.map_injective hcomp] + end Functor /-- A sieve induces a presheaf. -/ From 08ada333c4a954307d703701033a0d2fb7809e18 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Fri, 29 Nov 2024 14:47:10 +0000 Subject: [PATCH 417/829] chore: update Mathlib dependencies 2024-11-29 (#19605) 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 42df1c2b614e6..708733745b971 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "dfbe9387054e2972449de7bf346059d3609529fa", + "rev": "5de63200cb5bb3460e09225874a6e2fe6ca9da5d", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From cf80f345a26ac93fc5f71ff245ae7089fcc97adc Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Fri, 29 Nov 2024 15:48:03 +0000 Subject: [PATCH 418/829] chore: update Mathlib dependencies 2024-11-29 (#19608) 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 708733745b971..7e9d6679c79f6 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "5de63200cb5bb3460e09225874a6e2fe6ca9da5d", + "rev": "f46c0445cc86ad2bc973edf8c5b2bd88bd91913b", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From f9e900aae7c006d8e9a0f23f222ceb91e522f942 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 29 Nov 2024 19:27:58 +0000 Subject: [PATCH 419/829] chore(Filter/Prod): drop `Filter.prod`, use `SProd` instead (#18315) This way we can't accidentally use the underlying `Filter.prod` instead of the normal form. --- .../SpecialFunctions/NonIntegrable.lean | 2 +- Mathlib/Data/Analysis/Filter.lean | 2 +- Mathlib/Order/Filter/Basic.lean | 14 +++++++ Mathlib/Order/Filter/Curry.lean | 42 +++++++------------ Mathlib/Order/Filter/Defs.lean | 15 ++++--- Mathlib/Order/Filter/Finite.lean | 17 -------- Mathlib/Order/Filter/Prod.lean | 41 +++++++----------- .../Topology/Algebra/UniformGroup/Defs.lean | 16 ++----- Mathlib/Topology/Constructions.lean | 6 +-- Mathlib/Topology/Instances/TrivSqZeroExt.lean | 8 ++-- Mathlib/Topology/UniformSpace/Basic.lean | 3 +- .../UniformSpace/UniformConvergence.lean | 10 ++--- 12 files changed, 68 insertions(+), 108 deletions(-) diff --git a/Mathlib/Analysis/SpecialFunctions/NonIntegrable.lean b/Mathlib/Analysis/SpecialFunctions/NonIntegrable.lean index eb8ceabe6f887..8567fd824e7a8 100644 --- a/Mathlib/Analysis/SpecialFunctions/NonIntegrable.lean +++ b/Mathlib/Analysis/SpecialFunctions/NonIntegrable.lean @@ -58,7 +58,7 @@ theorem not_integrableOn_of_tendsto_norm_atTop_of_deriv_isBigO_filter_aux (∀ x ∈ s, ∀ y ∈ s, ∀ z ∈ [[x, y]], DifferentiableAt ℝ f z) ∧ ∀ x ∈ s, ∀ y ∈ s, ∀ z ∈ [[x, y]], ‖deriv f z‖ ≤ C * ‖g z‖ := by rcases hfg.exists_nonneg with ⟨C, C₀, hC⟩ - have h : ∀ᶠ x : ℝ × ℝ in l.prod l, + have h : ∀ᶠ x : ℝ × ℝ in l ×ˢ l, ∀ y ∈ [[x.1, x.2]], (DifferentiableAt ℝ f y ∧ ‖deriv f y‖ ≤ C * ‖g y‖) ∧ y ∈ k := (tendsto_fst.uIcc tendsto_snd).eventually ((hd.and hC.bound).and hl).smallSets rcases mem_prod_self_iff.1 h with ⟨s, hsl, hs⟩ diff --git a/Mathlib/Data/Analysis/Filter.lean b/Mathlib/Data/Analysis/Filter.lean index 46a34c12aec98..39e01fe198302 100644 --- a/Mathlib/Data/Analysis/Filter.lean +++ b/Mathlib/Data/Analysis/Filter.lean @@ -286,7 +286,7 @@ protected def iSup {f : α → Filter β} (F : ∀ i, (f i).Realizer) : (⨆ i, ⟨fun ⟨_, f⟩ i ↦ f i ⟨⟩, fun f ↦ ⟨(), fun i _ ↦ f i⟩, fun _ ↦ rfl, fun _ ↦ rfl⟩ /-- Construct a realizer for the product of filters -/ -protected def prod {f g : Filter α} (F : f.Realizer) (G : g.Realizer) : (f.prod g).Realizer := +protected def prod {f g : Filter α} (F : f.Realizer) (G : g.Realizer) : (f ×ˢ g).Realizer := (F.comap _).inf (G.comap _) theorem le_iff {f g : Filter α} (F : f.Realizer) (G : g.Realizer) : diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index 65c9595abbf40..ddfb081b4e76b 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -646,6 +646,10 @@ theorem eventually_or_distrib_right {f : Filter α} {p : α → Prop} {q : Prop} (∀ᶠ x in f, p x ∨ q) ↔ (∀ᶠ x in f, p x) ∨ q := by simp only [@or_comm _ q, eventually_or_distrib_left] +theorem eventually_imp_distrib_left {f : Filter α} {p : Prop} {q : α → Prop} : + (∀ᶠ x in f, p → q x) ↔ p → ∀ᶠ x in f, q x := by + simp only [imp_iff_not_or, eventually_or_distrib_left] + @[simp] theorem eventually_bot {p : α → Prop} : ∀ᶠ x in ⊥, p x := ⟨⟩ @@ -798,6 +802,16 @@ theorem eventually_imp_distrib_right {f : Filter α} {p : α → Prop} {q : Prop (∀ᶠ x in f, p x → q) ↔ (∃ᶠ x in f, p x) → q := by simp only [imp_iff_not_or, eventually_or_distrib_right, not_frequently] +@[simp] +theorem frequently_and_distrib_left {f : Filter α} {p : Prop} {q : α → Prop} : + (∃ᶠ x in f, p ∧ q x) ↔ p ∧ ∃ᶠ x in f, q x := by + simp only [Filter.Frequently, not_and, eventually_imp_distrib_left, Classical.not_imp] + +@[simp] +theorem frequently_and_distrib_right {f : Filter α} {p : α → Prop} {q : Prop} : + (∃ᶠ x in f, p x ∧ q) ↔ (∃ᶠ x in f, p x) ∧ q := by + simp only [@and_comm _ q, frequently_and_distrib_left] + @[simp] theorem frequently_bot {p : α → Prop} : ¬∃ᶠ x in ⊥, p x := by simp diff --git a/Mathlib/Order/Filter/Curry.lean b/Mathlib/Order/Filter/Curry.lean index 7f91d3cc76ef9..5d0fb17e49630 100644 --- a/Mathlib/Order/Filter/Curry.lean +++ b/Mathlib/Order/Filter/Curry.lean @@ -47,46 +47,34 @@ uniform convergence, curried filters, product filters namespace Filter -variable {α β γ : Type*} +variable {α β γ : Type*} {l : Filter α} {m : Filter β} {s : Set α} {t : Set β} -theorem eventually_curry_iff {f : Filter α} {g : Filter β} {p : α × β → Prop} : - (∀ᶠ x : α × β in f.curry g, p x) ↔ ∀ᶠ x : α in f, ∀ᶠ y : β in g, p (x, y) := +theorem eventually_curry_iff {p : α × β → Prop} : + (∀ᶠ x : α × β in l.curry m, p x) ↔ ∀ᶠ x : α in l, ∀ᶠ y : β in m, p (x, y) := Iff.rfl -theorem frequently_curry_iff {α β : Type*} {l : Filter α} {m : Filter β} +theorem frequently_curry_iff (p : (α × β) → Prop) : (∃ᶠ x in l.curry m, p x) ↔ ∃ᶠ x in l, ∃ᶠ y in m, p (x, y) := by simp_rw [Filter.Frequently, not_iff_not, not_not, eventually_curry_iff] -theorem mem_curry_iff {f : Filter α} {g : Filter β} {s : Set (α × β)} : - s ∈ f.curry g ↔ ∀ᶠ x : α in f, ∀ᶠ y : β in g, (x, y) ∈ s := Iff.rfl +theorem mem_curry_iff {s : Set (α × β)} : + s ∈ l.curry m ↔ ∀ᶠ x : α in l, ∀ᶠ y : β in m, (x, y) ∈ s := Iff.rfl -theorem curry_le_prod {f : Filter α} {g : Filter β} : f.curry g ≤ f.prod g := - fun _ => Eventually.curry +theorem curry_le_prod : l.curry m ≤ l ×ˢ m := fun _ => Eventually.curry theorem Tendsto.curry {f : α → β → γ} {la : Filter α} {lb : Filter β} {lc : Filter γ} (h : ∀ᶠ a in la, Tendsto (fun b : β => f a b) lb lc) : Tendsto (↿f) (la.curry lb) lc := fun _s hs => h.mono fun _a ha => ha hs -theorem frequently_curry_prod_iff {α β : Type*} {l : Filter α} {m : Filter β} - (s : Set α) (t : Set β) : (∃ᶠ x in l.curry m, x ∈ s ×ˢ t) ↔ sᶜ ∉ l ∧ tᶜ ∉ m := by - refine ⟨fun h => ?_, fun ⟨hs, ht⟩ => ?_⟩ - · exact frequently_prod_and.mp (Frequently.filter_mono h curry_le_prod) - rw [frequently_curry_iff] - exact Frequently.mono hs <| fun x hx => Frequently.mono ht (by simp [hx]) +theorem frequently_curry_prod_iff : + (∃ᶠ x in l.curry m, x ∈ s ×ˢ t) ↔ (∃ᶠ x in l, x ∈ s) ∧ ∃ᶠ y in m, y ∈ t := by + simp [frequently_curry_iff] -theorem prod_mem_curry {α β : Type*} {l : Filter α} {m : Filter β} {s : Set α} {t : Set β} - (hs : s ∈ l) (ht : t ∈ m) : s ×ˢ t ∈ l.curry m := - curry_le_prod <| prod_mem_prod hs ht - -theorem eventually_curry_prod_iff {α β : Type*} {l : Filter α} {m : Filter β} - [NeBot l] [NeBot m] (s : Set α) (t : Set β) : +theorem eventually_curry_prod_iff [NeBot l] [NeBot m] : (∀ᶠ x in l.curry m, x ∈ s ×ˢ t) ↔ s ∈ l ∧ t ∈ m := by - refine ⟨fun h => ⟨?_, ?_⟩, fun ⟨hs, ht⟩ => prod_mem_curry hs ht⟩ <;> - rw [eventually_curry_iff] at h - · apply mem_of_superset h - simp - rcases h.exists with ⟨_, hx⟩ - apply mem_of_superset hx - exact fun _ hy => hy.2 + simp [eventually_curry_iff] + +theorem prod_mem_curry (hs : s ∈ l) (ht : t ∈ m) : s ×ˢ t ∈ l.curry m := + curry_le_prod <| prod_mem_prod hs ht end Filter diff --git a/Mathlib/Order/Filter/Defs.lean b/Mathlib/Order/Filter/Defs.lean index d0a00823afb8a..4cf86091425d0 100644 --- a/Mathlib/Order/Filter/Defs.lean +++ b/Mathlib/Order/Filter/Defs.lean @@ -301,17 +301,20 @@ def comap (m : α → β) (f : Filter β) : Filter α where inter_sets := fun ⟨a', ha₁, ha₂⟩ ⟨b', hb₁, hb₂⟩ => ⟨a' ∩ b', inter_mem ha₁ hb₁, inter_subset_inter ha₂ hb₂⟩ -/-- Product of filters. This is the filter generated by cartesian products -of elements of the component filters. -/ -protected def prod (f : Filter α) (g : Filter β) : Filter (α × β) := - f.comap Prod.fst ⊓ g.comap Prod.snd - /-- Coproduct of filters. -/ protected def coprod (f : Filter α) (g : Filter β) : Filter (α × β) := f.comap Prod.fst ⊔ g.comap Prod.snd +/-- Product of filters. This is the filter generated by cartesian products +of elements of the component filters. -/ instance instSProd : SProd (Filter α) (Filter β) (Filter (α × β)) where - sprod := Filter.prod + sprod f g := f.comap Prod.fst ⊓ g.comap Prod.snd + +/-- Product of filters. This is the filter generated by cartesian products +of elements of the component filters. +Deprecated. Use `f ×ˢ g` instead. -/ +@[deprecated (since := "2024-11-29")] +protected def prod (f : Filter α) (g : Filter β) : Filter (α × β) := f ×ˢ g theorem prod_eq_inf (f : Filter α) (g : Filter β) : f ×ˢ g = f.comap Prod.fst ⊓ g.comap Prod.snd := rfl diff --git a/Mathlib/Order/Filter/Finite.lean b/Mathlib/Order/Filter/Finite.lean index 2c96231b9a89f..a10cfa503fa90 100644 --- a/Mathlib/Order/Filter/Finite.lean +++ b/Mathlib/Order/Filter/Finite.lean @@ -263,23 +263,6 @@ alias _root_.Finset.eventually_all := eventually_all_finset -- attribute [protected] Finset.eventually_all -theorem eventually_imp_distrib_left {f : Filter α} {p : Prop} {q : α → Prop} : - (∀ᶠ x in f, p → q x) ↔ p → ∀ᶠ x in f, q x := - eventually_all - - -/-! ### Frequently -/ - -@[simp] -theorem frequently_and_distrib_left {f : Filter α} {p : Prop} {q : α → Prop} : - (∃ᶠ x in f, p ∧ q x) ↔ p ∧ ∃ᶠ x in f, q x := by - simp only [Filter.Frequently, not_and, eventually_imp_distrib_left, Classical.not_imp] - -@[simp] -theorem frequently_and_distrib_right {f : Filter α} {p : α → Prop} {q : Prop} : - (∃ᶠ x in f, p x ∧ q) ↔ (∃ᶠ x in f, p x) ∧ q := by - simp only [@and_comm _ q, frequently_and_distrib_left] - /-! ### Relation “eventually equal” -/ diff --git a/Mathlib/Order/Filter/Prod.lean b/Mathlib/Order/Filter/Prod.lean index 5a1d1ab10faa6..0f0c40960a7ab 100644 --- a/Mathlib/Order/Filter/Prod.lean +++ b/Mathlib/Order/Filter/Prod.lean @@ -51,7 +51,6 @@ theorem prod_mem_prod (hs : s ∈ f) (ht : t ∈ g) : s ×ˢ t ∈ f ×ˢ g := theorem mem_prod_iff {s : Set (α × β)} {f : Filter α} {g : Filter β} : s ∈ f ×ˢ g ↔ ∃ t₁ ∈ f, ∃ t₂ ∈ g, t₁ ×ˢ t₂ ⊆ s := by - simp only [SProd.sprod, Filter.prod] constructor · rintro ⟨t₁, ⟨s₁, hs₁, hts₁⟩, t₂, ⟨s₂, hs₂, hts₂⟩, rfl⟩ exact ⟨s₁, hs₁, s₂, hs₂, fun p ⟨h, h'⟩ => ⟨hts₁ h, hts₂ h'⟩⟩ @@ -100,20 +99,16 @@ theorem comap_prodMap_prod (f : α → β) (g : γ → δ) (lb : Filter β) (ld simp [prod_eq_inf, comap_comap, Function.comp_def] theorem prod_top : f ×ˢ (⊤ : Filter β) = f.comap Prod.fst := by - dsimp only [SProd.sprod] - rw [Filter.prod, comap_top, inf_top_eq] + rw [prod_eq_inf, comap_top, inf_top_eq] theorem top_prod : (⊤ : Filter α) ×ˢ g = g.comap Prod.snd := by - dsimp only [SProd.sprod] - rw [Filter.prod, comap_top, top_inf_eq] + rw [prod_eq_inf, comap_top, top_inf_eq] theorem sup_prod (f₁ f₂ : Filter α) (g : Filter β) : (f₁ ⊔ f₂) ×ˢ g = (f₁ ×ˢ g) ⊔ (f₂ ×ˢ g) := by - dsimp only [SProd.sprod] - rw [Filter.prod, comap_sup, inf_sup_right, ← Filter.prod, ← Filter.prod] + simp only [prod_eq_inf, comap_sup, inf_sup_right] theorem prod_sup (f : Filter α) (g₁ g₂ : Filter β) : f ×ˢ (g₁ ⊔ g₂) = (f ×ˢ g₁) ⊔ (f ×ˢ g₂) := by - dsimp only [SProd.sprod] - rw [Filter.prod, comap_sup, inf_sup_left, ← Filter.prod, ← Filter.prod] + simp only [prod_eq_inf, comap_sup, inf_sup_left] theorem eventually_prod_iff {p : α × β → Prop} : (∀ᶠ x in f ×ˢ g, p x) ↔ @@ -203,15 +198,11 @@ theorem tendsto_diag : Tendsto (fun i => (i, i)) f (f ×ˢ f) := theorem prod_iInf_left [Nonempty ι] {f : ι → Filter α} {g : Filter β} : (⨅ i, f i) ×ˢ g = ⨅ i, f i ×ˢ g := by - dsimp only [SProd.sprod] - rw [Filter.prod, comap_iInf, iInf_inf] - simp only [Filter.prod, eq_self_iff_true] + simp only [prod_eq_inf, comap_iInf, iInf_inf] theorem prod_iInf_right [Nonempty ι] {f : Filter α} {g : ι → Filter β} : (f ×ˢ ⨅ i, g i) = ⨅ i, f ×ˢ g i := by - dsimp only [SProd.sprod] - rw [Filter.prod, comap_iInf, inf_iInf] - simp only [Filter.prod, eq_self_iff_true] + simp only [prod_eq_inf, comap_iInf, inf_iInf] @[mono, gcongr] theorem prod_mono {f₁ f₂ : Filter α} {g₁ g₂ : Filter β} (hf : f₁ ≤ f₂) (hg : g₁ ≤ g₂) : @@ -229,11 +220,10 @@ theorem prod_mono_right (f : Filter α) {g₁ g₂ : Filter β} (hf : g₁ ≤ g theorem prod_comap_comap_eq.{u, v, w, x} {α₁ : Type u} {α₂ : Type v} {β₁ : Type w} {β₂ : Type x} {f₁ : Filter α₁} {f₂ : Filter α₂} {m₁ : β₁ → α₁} {m₂ : β₂ → α₂} : comap m₁ f₁ ×ˢ comap m₂ f₂ = comap (fun p : β₁ × β₂ => (m₁ p.1, m₂ p.2)) (f₁ ×ˢ f₂) := by - simp only [SProd.sprod, Filter.prod, comap_comap, comap_inf, Function.comp_def] + simp only [prod_eq_inf, comap_comap, comap_inf, Function.comp_def] theorem prod_comm' : f ×ˢ g = comap Prod.swap (g ×ˢ f) := by - simp only [SProd.sprod, Filter.prod, comap_comap, Function.comp_def, inf_comm, Prod.swap, - comap_inf] + simp only [prod_eq_inf, comap_comap, Function.comp_def, inf_comm, Prod.swap, comap_inf] theorem prod_comm : f ×ˢ g = map (fun p : β × α => (p.2, p.1)) (g ×ˢ f) := by rw [prod_comm', ← map_swap_eq_comap_swap] @@ -280,12 +270,12 @@ theorem eventually_swap_iff {p : α × β → Prop} : theorem prod_assoc (f : Filter α) (g : Filter β) (h : Filter γ) : map (Equiv.prodAssoc α β γ) ((f ×ˢ g) ×ˢ h) = f ×ˢ (g ×ˢ h) := by - simp_rw [← comap_equiv_symm, SProd.sprod, Filter.prod, comap_inf, comap_comap, inf_assoc, + simp_rw [← comap_equiv_symm, prod_eq_inf, comap_inf, comap_comap, inf_assoc, Function.comp_def, Equiv.prodAssoc_symm_apply] theorem prod_assoc_symm (f : Filter α) (g : Filter β) (h : Filter γ) : map (Equiv.prodAssoc α β γ).symm (f ×ˢ (g ×ˢ h)) = (f ×ˢ g) ×ˢ h := by - simp_rw [map_equiv_symm, SProd.sprod, Filter.prod, comap_inf, comap_comap, inf_assoc, + simp_rw [map_equiv_symm, prod_eq_inf, comap_inf, comap_comap, inf_assoc, Function.comp_def, Equiv.prodAssoc_apply] theorem tendsto_prodAssoc {h : Filter γ} : @@ -300,7 +290,7 @@ theorem tendsto_prodAssoc_symm {h : Filter γ} : theorem map_swap4_prod {h : Filter γ} {k : Filter δ} : map (fun p : (α × β) × γ × δ => ((p.1.1, p.2.1), (p.1.2, p.2.2))) ((f ×ˢ g) ×ˢ (h ×ˢ k)) = (f ×ˢ h) ×ˢ (g ×ˢ k) := by - simp_rw [map_swap4_eq_comap, SProd.sprod, Filter.prod, comap_inf, comap_comap]; ac_rfl + simp_rw [map_swap4_eq_comap, prod_eq_inf, comap_inf, comap_comap]; ac_rfl theorem tendsto_swap4_prod {h : Filter γ} {k : Filter δ} : Tendsto (fun p : (α × β) × γ × δ => ((p.1.1, p.2.1), (p.1.2, p.2.2))) ((f ×ˢ g) ×ˢ (h ×ˢ k)) @@ -351,7 +341,7 @@ theorem prod_eq : f ×ˢ g = (f.map Prod.mk).seq g := f.map_prod id g theorem prod_inf_prod {f₁ f₂ : Filter α} {g₁ g₂ : Filter β} : (f₁ ×ˢ g₁) ⊓ (f₂ ×ˢ g₂) = (f₁ ⊓ f₂) ×ˢ (g₁ ⊓ g₂) := by - simp only [SProd.sprod, Filter.prod, comap_inf, inf_comm, inf_assoc, inf_left_comm] + simp only [prod_eq_inf, comap_inf, inf_comm, inf_assoc, inf_left_comm] theorem inf_prod {f₁ f₂ : Filter α} : (f₁ ⊓ f₂) ×ˢ g = (f₁ ×ˢ g) ⊓ (f₂ ×ˢ g) := by rw [prod_inf_prod, inf_idem] @@ -361,8 +351,7 @@ theorem prod_inf {g₁ g₂ : Filter β} : f ×ˢ (g₁ ⊓ g₂) = (f ×ˢ g₁ @[simp] theorem prod_principal_principal {s : Set α} {t : Set β} : 𝓟 s ×ˢ 𝓟 t = 𝓟 (s ×ˢ t) := by - simp only [SProd.sprod, Filter.prod, comap_principal, principal_eq_iff_eq, comap_principal, - inf_principal]; rfl + simp only [prod_eq_inf, comap_principal, principal_eq_iff_eq, comap_principal, inf_principal]; rfl @[simp] theorem pure_prod {a : α} {f : Filter β} : pure a ×ˢ f = map (Prod.mk a) f := by @@ -413,9 +402,7 @@ theorem tendsto_prod_iff {f : α × β → γ} {x : Filter α} {y : Filter β} { theorem tendsto_prod_iff' {g' : Filter γ} {s : α → β × γ} : Tendsto s f (g ×ˢ g') ↔ Tendsto (fun n => (s n).1) f g ∧ Tendsto (fun n => (s n).2) f g' := by - dsimp only [SProd.sprod] - unfold Filter.prod - simp only [tendsto_inf, tendsto_comap_iff, Function.comp_def] + simp only [prod_eq_inf, tendsto_inf, tendsto_comap_iff, Function.comp_def] theorem le_prod {f : Filter (α × β)} {g : Filter α} {g' : Filter β} : (f ≤ g ×ˢ g') ↔ Tendsto Prod.fst f g ∧ Tendsto Prod.snd f g' := diff --git a/Mathlib/Topology/Algebra/UniformGroup/Defs.lean b/Mathlib/Topology/Algebra/UniformGroup/Defs.lean index 3455d9f14d056..5b6a3aeeb9ff2 100644 --- a/Mathlib/Topology/Algebra/UniformGroup/Defs.lean +++ b/Mathlib/Topology/Algebra/UniformGroup/Defs.lean @@ -406,19 +406,11 @@ variable {G} @[to_additive] -- Porting note: renamed theorem to conform to naming convention theorem comm_topologicalGroup_is_uniform : UniformGroup G := by - have : - Tendsto - ((fun p : G × G => p.1 / p.2) ∘ fun p : (G × G) × G × G => (p.1.2 / p.1.1, p.2.2 / p.2.1)) - (comap (fun p : (G × G) × G × G => (p.1.2 / p.1.1, p.2.2 / p.2.1)) ((𝓝 1).prod (𝓝 1))) - (𝓝 (1 / 1)) := - (tendsto_fst.div' tendsto_snd).comp tendsto_comap constructor - rw [UniformContinuous, uniformity_prod_eq_prod, tendsto_map'_iff, uniformity_eq_comap_nhds_one' G, - tendsto_comap_iff, prod_comap_comap_eq] - simp only [Function.comp_def, div_eq_mul_inv, mul_inv_rev, inv_inv, mul_comm, mul_left_comm] at * - simp only [inv_one, mul_one, ← mul_assoc] at this - simp_rw [← mul_assoc, mul_comm] - assumption + simp only [UniformContinuous, uniformity_prod_eq_prod, uniformity_eq_comap_nhds_one', + tendsto_comap_iff, tendsto_map'_iff, prod_comap_comap_eq, Function.comp_def, + div_div_div_comm _ (Prod.snd (Prod.snd _)), ← nhds_prod_eq, Prod.mk_one_one] + exact (continuous_div'.tendsto' 1 1 (div_one 1)).comp tendsto_comap open Set diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index 280967dfdc29d..97a7797155a00 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -477,11 +477,9 @@ theorem IsOpen.prod {s : Set X} {t : Set Y} (hs : IsOpen s) (ht : IsOpen t) : Is -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11215): TODO: Lean fails to find `t₁` and `t₂` by unification theorem nhds_prod_eq {x : X} {y : Y} : 𝓝 (x, y) = 𝓝 x ×ˢ 𝓝 y := by - dsimp only [SProd.sprod] - rw [Filter.prod, instTopologicalSpaceProd, nhds_inf (t₁ := TopologicalSpace.induced Prod.fst _) + rw [prod_eq_inf, instTopologicalSpaceProd, nhds_inf (t₁ := TopologicalSpace.induced Prod.fst _) (t₂ := TopologicalSpace.induced Prod.snd _), nhds_induced, nhds_induced] --- Porting note: moved from `Topology.ContinuousOn` theorem nhdsWithin_prod_eq (x : X) (y : Y) (s : Set X) (t : Set Y) : 𝓝[s ×ˢ t] (x, y) = 𝓝[s] x ×ˢ 𝓝[t] y := by simp only [nhdsWithin, nhds_prod_eq, ← prod_inf_prod, prod_principal_principal] @@ -504,7 +502,6 @@ theorem mem_nhdsWithin_prod_iff {x : X} {y : Y} {s : Set (X × Y)} {tx : Set X} s ∈ 𝓝[tx ×ˢ ty] (x, y) ↔ ∃ u ∈ 𝓝[tx] x, ∃ v ∈ 𝓝[ty] y, u ×ˢ v ⊆ s := by rw [nhdsWithin_prod_eq, mem_prod_iff] --- Porting note: moved up theorem Filter.HasBasis.prod_nhds {ιX ιY : Type*} {px : ιX → Prop} {py : ιY → Prop} {sx : ιX → Set X} {sy : ιY → Set Y} {x : X} {y : Y} (hx : (𝓝 x).HasBasis px sx) (hy : (𝓝 y).HasBasis py sy) : @@ -512,7 +509,6 @@ theorem Filter.HasBasis.prod_nhds {ιX ιY : Type*} {px : ιX → Prop} {py : ι rw [nhds_prod_eq] exact hx.prod hy --- Porting note: moved up theorem Filter.HasBasis.prod_nhds' {ιX ιY : Type*} {pX : ιX → Prop} {pY : ιY → Prop} {sx : ιX → Set X} {sy : ιY → Set Y} {p : X × Y} (hx : (𝓝 p.1).HasBasis pX sx) (hy : (𝓝 p.2).HasBasis pY sy) : diff --git a/Mathlib/Topology/Instances/TrivSqZeroExt.lean b/Mathlib/Topology/Instances/TrivSqZeroExt.lean index b8a913a8e6a65..5a6add1956758 100644 --- a/Mathlib/Topology/Instances/TrivSqZeroExt.lean +++ b/Mathlib/Topology/Instances/TrivSqZeroExt.lean @@ -42,14 +42,12 @@ instance instTopologicalSpace : TopologicalSpace (tsze R M) := instance [T2Space R] [T2Space M] : T2Space (tsze R M) := Prod.t2Space -theorem nhds_def (x : tsze R M) : 𝓝 x = (𝓝 x.fst).prod (𝓝 x.snd) := by - cases x using Prod.rec - exact nhds_prod_eq +theorem nhds_def (x : tsze R M) : 𝓝 x = 𝓝 x.fst ×ˢ 𝓝 x.snd := nhds_prod_eq -theorem nhds_inl [Zero M] (x : R) : 𝓝 (inl x : tsze R M) = (𝓝 x).prod (𝓝 0) := +theorem nhds_inl [Zero M] (x : R) : 𝓝 (inl x : tsze R M) = 𝓝 x ×ˢ 𝓝 0 := nhds_def _ -theorem nhds_inr [Zero R] (m : M) : 𝓝 (inr m : tsze R M) = (𝓝 0).prod (𝓝 m) := +theorem nhds_inr [Zero R] (m : M) : 𝓝 (inr m : tsze R M) = 𝓝 0 ×ˢ 𝓝 m := nhds_def _ nonrec theorem continuous_fst : Continuous (fst : tsze R M → R) := diff --git a/Mathlib/Topology/UniformSpace/Basic.lean b/Mathlib/Topology/UniformSpace/Basic.lean index 5534b8ebfe86d..9f28773dd89af 100644 --- a/Mathlib/Topology/UniformSpace/Basic.lean +++ b/Mathlib/Topology/UniformSpace/Basic.lean @@ -1393,8 +1393,7 @@ instance [UniformSpace α] [IsCountablyGenerated (𝓤 α)] theorem uniformity_prod_eq_comap_prod [UniformSpace α] [UniformSpace β] : 𝓤 (α × β) = comap (fun p : (α × β) × α × β => ((p.1.1, p.2.1), (p.1.2, p.2.2))) (𝓤 α ×ˢ 𝓤 β) := by - dsimp [SProd.sprod] - rw [uniformity_prod, Filter.prod, Filter.comap_inf, Filter.comap_comap, Filter.comap_comap]; rfl + simp_rw [uniformity_prod, prod_eq_inf, Filter.comap_inf, Filter.comap_comap, Function.comp_def] theorem uniformity_prod_eq_prod [UniformSpace α] [UniformSpace β] : 𝓤 (α × β) = map (fun p : (α × α) × β × β => ((p.1.1, p.2.1), (p.1.2, p.2.2))) (𝓤 α ×ˢ 𝓤 β) := by diff --git a/Mathlib/Topology/UniformSpace/UniformConvergence.lean b/Mathlib/Topology/UniformSpace/UniformConvergence.lean index 48ea1c9d19316..d6bb495deea18 100644 --- a/Mathlib/Topology/UniformSpace/UniformConvergence.lean +++ b/Mathlib/Topology/UniformSpace/UniformConvergence.lean @@ -273,10 +273,11 @@ theorem TendstoUniformlyOnFilter.prod {ι' β' : Type*} [UniformSpace β'] {F' : (p ×ˢ q) p' := fun u hu => ((h.prod_map h') u hu).diag_of_prod_right -theorem TendstoUniformlyOn.prod {ι' β' : Type*} [UniformSpace β'] {F' : ι' → α → β'} {f' : α → β'} - {p' : Filter ι'} (h : TendstoUniformlyOn F f p s) (h' : TendstoUniformlyOn F' f' p' s) : +protected theorem TendstoUniformlyOn.prod {ι' β' : Type*} [UniformSpace β'] + {F' : ι' → α → β'} {f' : α → β'} {p' : Filter ι'} + (h : TendstoUniformlyOn F f p s) (h' : TendstoUniformlyOn F' f' p' s) : TendstoUniformlyOn (fun (i : ι × ι') a => (F i.1 a, F' i.2 a)) (fun a => (f a, f' a)) - (p.prod p') s := + (p ×ˢ p') s := (congr_arg _ s.inter_self).mp ((h.prod_map h').comp fun a => (a, a)) theorem TendstoUniformly.prod {ι' β' : Type*} [UniformSpace β'] {F' : ι' → α → β'} {f' : α → β'} @@ -333,8 +334,7 @@ theorem UniformContinuousOn.tendstoUniformlyOn [UniformSpace α] [UniformSpace set φ := fun q : α × β => ((x, q.2), q) rw [tendstoUniformlyOn_iff_tendsto] change Tendsto (Prod.map (↿F) ↿F ∘ φ) (𝓝[U] x ×ˢ 𝓟 V) (𝓤 γ) - simp only [nhdsWithin, SProd.sprod, Filter.prod, comap_inf, inf_assoc, comap_principal, - inf_principal] + simp only [nhdsWithin, Filter.prod_eq_inf, comap_inf, inf_assoc, comap_principal, inf_principal] refine hF.comp (Tendsto.inf ?_ <| tendsto_principal_principal.2 fun x hx => ⟨⟨hU, hx.2⟩, hx⟩) simp only [uniformity_prod_eq_comap_prod, tendsto_comap_iff, (· ∘ ·), nhds_eq_comap_uniformity, comap_comap] From 67afcb68326509120b289a2848cbba4162b616a9 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Fri, 29 Nov 2024 21:45:10 +0000 Subject: [PATCH 420/829] feat: update the `ContDiff` definition to integrate analytic functions in the hierarchy (#17152) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We modify the definition of `ContDiff` so that the smoothness exponent belongs to `WithTop (ℕ∞)`, where top smoothness (denoted with `ω` in the `ContDiff` scope) means analyticity (together with analyticity of all the iterated derivatives). This will make it possible to write theorems like: consider a function which is `C^2` over the reals or complexes, or analytic over a general field. Then... Many theorems in calculus hold under these assumptions (and it's the way things are written in Bourbaki), so having this possibility is a necessary improvement unless one wants to duplicate everything. The PR is mostly not changing the API: theorems that were proved for `C^n` functions with `n ∈ ℕ or n = ∞` are now proved for `n ∈ ℕ or n = ∞ or n = ω` (unless they are not true for analytic functions, but it turns out in most cases things are fine). Of course, results have to be added in the files `ContDiff.Def` and `ContDiff.Basic` to deal with analytic functions. I know that the PR is huge and hard to review. I have tried to separate all the other results in other PRs, but once one changes the basic definition one has to fix everything downstream to get mathlib compiling, so I can not make it shorter... Co-authored-by: Yury G. Kudryashov --- Mathlib.lean | 1 - Mathlib/Analysis/Analytic/OfScalars.lean | 3 +- .../Analysis/Calculus/ContDiff/Analytic.lean | 55 -- Mathlib/Analysis/Calculus/ContDiff/Basic.lean | 507 ++++++------ .../Calculus/ContDiff/CPolynomial.lean | 15 +- Mathlib/Analysis/Calculus/ContDiff/Defs.lean | 745 ++++++++++++------ .../Calculus/ContDiff/FiniteDimension.lean | 23 +- .../Analysis/Calculus/ContDiff/WithLp.lean | 10 +- .../Analysis/Calculus/FDeriv/Symmetric.lean | 17 +- .../Analysis/Calculus/IteratedDeriv/Defs.lean | 21 +- Mathlib/Analysis/Calculus/Taylor.lean | 3 +- Mathlib/Analysis/Complex/CauchyIntegral.lean | 7 +- Mathlib/Analysis/Convolution.lean | 6 +- .../Analysis/Distribution/SchwartzSpace.lean | 6 +- .../Analysis/InnerProductSpace/Calculus.lean | 10 +- .../NormedSpace/Multilinear/Basic.lean | 4 +- .../SpecialFunctions/Complex/Analytic.lean | 8 +- .../SpecialFunctions/Complex/LogDeriv.lean | 2 +- .../Analysis/SpecialFunctions/ExpDeriv.lean | 1 - .../Analysis/SpecialFunctions/Log/Basic.lean | 15 + .../Analysis/SpecialFunctions/Log/Deriv.lean | 28 +- .../Analysis/SpecialFunctions/Pow/Deriv.lean | 7 +- .../SpecialFunctions/SmoothTransition.lean | 3 +- Mathlib/Analysis/SpecialFunctions/Sqrt.lean | 6 +- .../Trigonometric/ArctanDeriv.lean | 4 +- .../Geometry/Manifold/Algebra/LieGroup.lean | 11 +- .../Geometry/Manifold/AnalyticManifold.lean | 13 +- Mathlib/Geometry/Manifold/ContMDiff/Defs.lean | 6 +- .../Manifold/Sheaf/LocallyRingedSpace.lean | 2 +- Mathlib/Tactic/FunProp/ContDiff.lean | 2 +- 30 files changed, 910 insertions(+), 631 deletions(-) delete mode 100644 Mathlib/Analysis/Calculus/ContDiff/Analytic.lean diff --git a/Mathlib.lean b/Mathlib.lean index d8c2ef98c4eed..5ce336f55f28b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1128,7 +1128,6 @@ import Mathlib.Analysis.Calculus.BumpFunction.InnerProduct import Mathlib.Analysis.Calculus.BumpFunction.Normed import Mathlib.Analysis.Calculus.Conformal.InnerProduct import Mathlib.Analysis.Calculus.Conformal.NormedSpace -import Mathlib.Analysis.Calculus.ContDiff.Analytic import Mathlib.Analysis.Calculus.ContDiff.Basic import Mathlib.Analysis.Calculus.ContDiff.Bounds import Mathlib.Analysis.Calculus.ContDiff.CPolynomial diff --git a/Mathlib/Analysis/Analytic/OfScalars.lean b/Mathlib/Analysis/Analytic/OfScalars.lean index 688d6468fb486..08114dc0ef413 100644 --- a/Mathlib/Analysis/Analytic/OfScalars.lean +++ b/Mathlib/Analysis/Analytic/OfScalars.lean @@ -62,7 +62,8 @@ variable (𝕜) in theorem ofScalars_series_injective [Nontrivial E] : Function.Injective (ofScalars E (𝕜 := 𝕜)) := by intro _ _ refine Function.mtr fun h ↦ ?_ - simp_rw [FormalMultilinearSeries.ext_iff, ofScalars, ContinuousMultilinearMap.ext_iff, smul_apply] + simp_rw [FormalMultilinearSeries.ext_iff, ofScalars, ContinuousMultilinearMap.ext_iff, + ContinuousMultilinearMap.smul_apply] push_neg obtain ⟨n, hn⟩ := Function.ne_iff.1 h refine ⟨n, fun _ ↦ 1, ?_⟩ diff --git a/Mathlib/Analysis/Calculus/ContDiff/Analytic.lean b/Mathlib/Analysis/Calculus/ContDiff/Analytic.lean deleted file mode 100644 index 8ba3cda66b843..0000000000000 --- a/Mathlib/Analysis/Calculus/ContDiff/Analytic.lean +++ /dev/null @@ -1,55 +0,0 @@ -/- -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.Analysis.Calculus.ContDiff.Defs -import Mathlib.Analysis.Calculus.FDeriv.Analytic - -/-! -# Analytic functions are `C^∞`. --/ - -open Set ContDiff - -variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] - {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] - {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] - {f : E → F} {s : Set E} {x : E} {n : WithTop ℕ∞} - -/-- An analytic function is infinitely differentiable. -/ -protected theorem AnalyticOnNhd.contDiffOn [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f s) : - ContDiffOn 𝕜 n f s := by - let t := { x | AnalyticAt 𝕜 f x } - suffices ContDiffOn 𝕜 ω f t from (this.of_le le_top).mono h - rw [← contDiffOn_infty_iff_contDiffOn_omega] - have H : AnalyticOnNhd 𝕜 f t := fun _x hx ↦ hx - have t_open : IsOpen t := isOpen_analyticAt 𝕜 f - exact contDiffOn_of_continuousOn_differentiableOn - (fun m _ ↦ (H.iteratedFDeriv m).continuousOn.congr - fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) - (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) : - 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 - exact hf.contDiffOn.contDiffAt hs - -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 - -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 diff --git a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean index ad3ed57edf02c..91cf96657a498 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean @@ -70,21 +70,13 @@ theorem iteratedFDeriv_zero_fun {n : ℕ} : (iteratedFDeriv 𝕜 n fun _ : E ↦ funext fun x ↦ by simpa [← iteratedFDerivWithin_univ] using iteratedFDerivWithin_zero_fun uniqueDiffOn_univ (mem_univ x) -theorem contDiff_zero_fun : ContDiff 𝕜 n fun _ : E => (0 : F) := by - suffices ContDiff 𝕜 ω (fun _ : E => (0 : F)) from this.of_le le_top - rw [← contDiff_infty_iff_contDiff_omega] - apply contDiff_of_differentiable_iteratedFDeriv fun m _ ↦ ?_ - rw [iteratedFDeriv_zero_fun] - exact differentiable_const (0 : E[×m]→L[𝕜] F) +theorem contDiff_zero_fun : ContDiff 𝕜 n fun _ : E => (0 : F) := + analyticOnNhd_const.contDiff /-- Constants are `C^∞`. -/ -theorem contDiff_const {c : F} : ContDiff 𝕜 n fun _ : E => c := by - suffices h : ContDiff 𝕜 ω fun _ : E => c from h.of_le le_top - rw [← contDiff_infty_iff_contDiff_omega, contDiff_top_iff_fderiv] - refine ⟨differentiable_const c, ?_⟩ - rw [fderiv_const] - exact contDiff_zero_fun +theorem contDiff_const {c : F} : ContDiff 𝕜 n fun _ : E => c := + analyticOnNhd_const.contDiff theorem contDiffOn_const {c : F} {s : Set E} : ContDiffOn 𝕜 n (fun _ : E => c) s := contDiff_const.contDiffOn @@ -141,14 +133,10 @@ theorem contDiffWithinAt_singleton : ContDiffWithinAt 𝕜 n f {x} x := /-! ### Smoothness of linear functions -/ -/-- Unbundled bounded linear functions are `C^∞`. +/-- Unbundled bounded linear functions are `C^n`. -/ -theorem IsBoundedLinearMap.contDiff (hf : IsBoundedLinearMap 𝕜 f) : ContDiff 𝕜 n f := by - suffices h : ContDiff 𝕜 ω f from h.of_le le_top - rw [← contDiff_infty_iff_contDiff_omega, contDiff_top_iff_fderiv] - refine ⟨hf.differentiable, ?_⟩ - simp_rw [hf.fderiv] - exact contDiff_const +theorem IsBoundedLinearMap.contDiff (hf : IsBoundedLinearMap 𝕜 f) : ContDiff 𝕜 n f := + (ContinuousLinearMap.analyticOnNhd hf.toContinuousLinearMap univ).contDiff theorem ContinuousLinearMap.contDiff (f : E →L[𝕜] F) : ContDiff 𝕜 n f := f.isBoundedLinearMap.contDiff @@ -162,7 +150,7 @@ theorem LinearIsometry.contDiff (f : E →ₗᵢ[𝕜] F) : ContDiff 𝕜 n f := theorem LinearIsometryEquiv.contDiff (f : E ≃ₗᵢ[𝕜] F) : ContDiff 𝕜 n f := (f : E →L[𝕜] F).contDiff -/-- The identity is `C^∞`. +/-- The identity is `C^n`. -/ theorem contDiff_id : ContDiff 𝕜 n (id : E → E) := IsBoundedLinearMap.id.contDiff @@ -176,14 +164,10 @@ theorem contDiffAt_id {x} : ContDiffAt 𝕜 n (id : E → E) x := theorem contDiffOn_id {s} : ContDiffOn 𝕜 n (id : E → E) s := contDiff_id.contDiffOn -/-- Bilinear functions are `C^∞`. +/-- Bilinear functions are `C^n`. -/ -theorem IsBoundedBilinearMap.contDiff (hb : IsBoundedBilinearMap 𝕜 b) : ContDiff 𝕜 n b := by - suffices h : ContDiff 𝕜 ω b from h.of_le le_top - rw [← contDiff_infty_iff_contDiff_omega, contDiff_top_iff_fderiv] - refine ⟨hb.differentiable, ?_⟩ - simp only [hb.fderiv] - exact hb.isBoundedLinearMap_deriv.contDiff +theorem IsBoundedBilinearMap.contDiff (hb : IsBoundedBilinearMap 𝕜 b) : ContDiff 𝕜 n b := + (hb.toContinuousLinearMap.analyticOnNhd_bilinear _).contDiff /-- If `f` admits a Taylor series `p` in a set `s`, and `g` is linear, then `g ∘ f` admits a Taylor series whose `k`-th term is given by `g ∘ (p k)`. -/ @@ -199,9 +183,20 @@ theorem HasFTaylorSeriesUpToOn.continuousLinearMap_comp {n : WithTop ℕ∞} (g /-- Composition by continuous linear maps on the left preserves `C^n` functions in a domain at a point. -/ theorem ContDiffWithinAt.continuousLinearMap_comp (g : F →L[𝕜] G) - (hf : ContDiffWithinAt 𝕜 n f s x) : ContDiffWithinAt 𝕜 n (g ∘ f) s x := fun m hm ↦ by - rcases hf m hm with ⟨u, hu, p, hp⟩ - exact ⟨u, hu, _, hp.continuousLinearMap_comp g⟩ + (hf : ContDiffWithinAt 𝕜 n f s x) : ContDiffWithinAt 𝕜 n (g ∘ f) s x := by + match n with + | ω => + obtain ⟨u, hu, p, hp, h'p⟩ := hf + refine ⟨u, hu, _, hp.continuousLinearMap_comp g, fun i ↦ ?_⟩ + change AnalyticOn 𝕜 + (fun x ↦ (ContinuousLinearMap.compContinuousMultilinearMapL 𝕜 + (fun _ : Fin i ↦ E) F G g) (p x i)) u + apply AnalyticOnNhd.comp_analyticOn _ (h'p i) (Set.mapsTo_univ _ _) + exact ContinuousLinearMap.analyticOnNhd _ _ + | (n : ℕ∞) => + intro m hm + rcases hf m hm with ⟨u, hu, p, hp⟩ + exact ⟨u, hu, _, hp.continuousLinearMap_comp g⟩ /-- Composition by continuous linear maps on the left preserves `C^n` functions in a domain at a point. -/ @@ -224,8 +219,8 @@ theorem ContinuousLinearMap.iteratedFDerivWithin_comp_left {f : E → F} (g : F (hf : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {i : ℕ} (hi : i ≤ n) : iteratedFDerivWithin 𝕜 i (g ∘ f) s x = g.compContinuousMultilinearMap (iteratedFDerivWithin 𝕜 i f s x) := - (((hf.ftaylorSeriesWithin hs).continuousLinearMap_comp g).eq_iteratedFDerivWithin_of_uniqueDiffOn - (mod_cast hi) hs hx).symm + ((((hf.of_le hi).ftaylorSeriesWithin hs).continuousLinearMap_comp + g).eq_iteratedFDerivWithin_of_uniqueDiffOn le_rfl hs hx).symm /-- The iterated derivative of the composition with a linear map on the left is obtained by applying the linear map to the iterated derivative. -/ @@ -326,7 +321,7 @@ theorem ContinuousLinearEquiv.comp_contDiff_iff (e : F ≃L[𝕜] G) : /-- If `f` admits a Taylor series `p` in a set `s`, and `g` is linear, then `f ∘ g` admits a Taylor series in `g ⁻¹' s`, whose `k`-th term is given by `p k (g v₁, ..., g vₖ)` . -/ -theorem HasFTaylorSeriesUpToOn.compContinuousLinearMap {n : WithTop ℕ∞} +theorem HasFTaylorSeriesUpToOn.compContinuousLinearMap (hf : HasFTaylorSeriesUpToOn n f p s) (g : G →L[𝕜] E) : HasFTaylorSeriesUpToOn n (f ∘ g) (fun x k => (p (g x) k).compContinuousLinearMap fun _ => g) (g ⁻¹' s) := by @@ -353,11 +348,24 @@ theorem HasFTaylorSeriesUpToOn.compContinuousLinearMap {n : WithTop ℕ∞} a domain. -/ theorem ContDiffWithinAt.comp_continuousLinearMap {x : G} (g : G →L[𝕜] E) (hf : ContDiffWithinAt 𝕜 n f s (g x)) : ContDiffWithinAt 𝕜 n (f ∘ g) (g ⁻¹' s) x := by - intro m hm - rcases hf m hm with ⟨u, hu, p, hp⟩ - refine ⟨g ⁻¹' u, ?_, _, hp.compContinuousLinearMap g⟩ - refine g.continuous.continuousWithinAt.tendsto_nhdsWithin ?_ hu - exact (mapsTo_singleton.2 <| mem_singleton _).union_union (mapsTo_preimage _ _) + match n with + | ω => + obtain ⟨u, hu, p, hp, h'p⟩ := hf + refine ⟨g ⁻¹' u, ?_, _, hp.compContinuousLinearMap g, ?_⟩ + · refine g.continuous.continuousWithinAt.tendsto_nhdsWithin ?_ hu + exact (mapsTo_singleton.2 <| mem_singleton _).union_union (mapsTo_preimage _ _) + · intro i + change AnalyticOn 𝕜 (fun x ↦ + ContinuousMultilinearMap.compContinuousLinearMapL (fun _ ↦ g) (p (g x) i)) (⇑g ⁻¹' u) + apply AnalyticOn.comp _ _ (Set.mapsTo_univ _ _) + · exact ContinuousLinearEquiv.analyticOn _ _ + · exact (h'p i).comp (g.analyticOn _) (mapsTo_preimage _ _) + | (n : ℕ∞) => + intro m hm + rcases hf m hm with ⟨u, hu, p, hp⟩ + refine ⟨g ⁻¹' u, ?_, _, hp.compContinuousLinearMap g⟩ + refine g.continuous.continuousWithinAt.tendsto_nhdsWithin ?_ hu + exact (mapsTo_singleton.2 <| mem_singleton _).union_union (mapsTo_preimage _ _) /-- Composition by continuous linear maps on the right preserves `C^n` functions on domains. -/ theorem ContDiffOn.comp_continuousLinearMap (hf : ContDiffOn 𝕜 n f s) (g : G →L[𝕜] E) : @@ -375,8 +383,8 @@ theorem ContinuousLinearMap.iteratedFDerivWithin_comp_right {f : E → F} (g : G (hx : g x ∈ s) {i : ℕ} (hi : i ≤ n) : iteratedFDerivWithin 𝕜 i (f ∘ g) (g ⁻¹' s) x = (iteratedFDerivWithin 𝕜 i f s (g x)).compContinuousLinearMap fun _ => g := - (((hf.ftaylorSeriesWithin hs).compContinuousLinearMap g).eq_iteratedFDerivWithin_of_uniqueDiffOn - (mod_cast hi) h's hx).symm + ((((hf.of_le hi).ftaylorSeriesWithin hs).compContinuousLinearMap + g).eq_iteratedFDerivWithin_of_uniqueDiffOn le_rfl h's hx).symm /-- The iterated derivative within a set of the composition with a linear equiv on the right is obtained by composing the iterated derivative with the linear equiv. -/ @@ -479,12 +487,24 @@ theorem HasFTaylorSeriesUpToOn.prod {n : WithTop ℕ∞} /-- The cartesian product of `C^n` functions at a point in a domain is `C^n`. -/ theorem ContDiffWithinAt.prod {s : Set E} {f : E → F} {g : E → G} (hf : ContDiffWithinAt 𝕜 n f s x) (hg : ContDiffWithinAt 𝕜 n g s x) : ContDiffWithinAt 𝕜 n (fun x : E => (f x, g x)) s x := by - intro m hm - rcases hf m hm with ⟨u, hu, p, hp⟩ - rcases hg m hm with ⟨v, hv, q, hq⟩ - exact - ⟨u ∩ v, Filter.inter_mem hu hv, _, - (hp.mono inter_subset_left).prod (hq.mono inter_subset_right)⟩ + match n with + | ω => + obtain ⟨u, hu, p, hp, h'p⟩ := hf + obtain ⟨v, hv, q, hq, h'q⟩ := hg + refine ⟨u ∩ v, Filter.inter_mem hu hv, _, + (hp.mono inter_subset_left).prod (hq.mono inter_subset_right), fun i ↦ ?_⟩ + change AnalyticOn 𝕜 (fun x ↦ ContinuousMultilinearMap.prodL _ _ _ _ (p x i, q x i)) + (u ∩ v) + apply AnalyticOnNhd.comp_analyticOn (LinearIsometryEquiv.analyticOnNhd _ _) _ + (Set.mapsTo_univ _ _) + exact ((h'p i).mono inter_subset_left).prod ((h'q i).mono inter_subset_right) + | (n : ℕ∞) => + intro m hm + rcases hf m hm with ⟨u, hu, p, hp⟩ + rcases hg m hm with ⟨v, hv, q, hq⟩ + exact + ⟨u ∩ v, Filter.inter_mem hu hv, _, + (hp.mono inter_subset_left).prod (hq.mono inter_subset_right)⟩ /-- The cartesian product of `C^n` functions on domains is `C^n`. -/ theorem ContDiffOn.prod {s : Set E} {f : E → F} {g : E → G} (hf : ContDiffOn 𝕜 n f s) @@ -514,6 +534,7 @@ multiplication is the application of a bilinear map (which is `C^∞`, and there `x ↦ (Dg(f x), Df x)`. As the composition of two `C^n` maps, it is again `C^n`, and we are done. There are two difficulties in this proof. + The first one is that it is an induction over all Banach spaces. In Lean, this is only possible if they belong to a fixed universe. One could formalize this by first proving the statement in this case, and then extending the result to general universes @@ -522,6 +543,7 @@ by embedding all the spaces we consider in a common universe through `ULift`. The second one is that it does not work cleanly for analytic maps: for this case, we need to exhibit a whole sequence of derivatives which are all analytic, not just finitely many of them, so an induction is never enough at a finite step. + Both these difficulties can be overcome with some cost. However, we choose a different path: we write down an explicit formula for the `n`-th derivative of `g ∘ f` in terms of derivatives of `g` and `f` (this is the formula of Faa-Di Bruno) and use this formula to get a suitable Taylor @@ -534,19 +556,44 @@ essentially trivial. theorem ContDiffWithinAt.comp {s : Set E} {t : Set F} {g : F → G} {f : E → F} (x : E) (hg : ContDiffWithinAt 𝕜 n g t (f x)) (hf : ContDiffWithinAt 𝕜 n f s x) (st : MapsTo f s t) : ContDiffWithinAt 𝕜 n (g ∘ f) s x := by - intro m hm - rcases hf m hm with ⟨u, hu, p, hp⟩ - rcases hg m hm with ⟨v, hv, q, hq⟩ - let w := insert x s ∩ (u ∩ f ⁻¹' v) - have wv : w ⊆ f ⁻¹' v := inter_subset_right.trans inter_subset_right - have wu : w ⊆ u := inter_subset_right.trans inter_subset_left - refine ⟨w, ?_, fun y ↦ (q (f y)).taylorComp (p y), hq.comp (hp.mono wu) wv⟩ - apply inter_mem self_mem_nhdsWithin (inter_mem hu ?_) - apply (continuousWithinAt_insert_self.2 hf.continuousWithinAt).preimage_mem_nhdsWithin' - apply nhdsWithin_mono _ _ hv - simp only [image_insert_eq] - apply insert_subset_insert - exact image_subset_iff.mpr st + match n with + | ω => + have h'f : ContDiffWithinAt 𝕜 ω f s x := hf + obtain ⟨u, hu, p, hp, h'p⟩ := h'f + obtain ⟨v, hv, q, hq, h'q⟩ := hg + let w := insert x s ∩ (u ∩ f ⁻¹' v) + have wv : w ⊆ f ⁻¹' v := fun y hy => hy.2.2 + have wu : w ⊆ u := fun y hy => hy.2.1 + refine ⟨w, ?_, fun y ↦ (q (f y)).taylorComp (p y), hq.comp (hp.mono wu) wv, ?_⟩ + · apply inter_mem self_mem_nhdsWithin (inter_mem hu ?_) + apply (continuousWithinAt_insert_self.2 hf.continuousWithinAt).preimage_mem_nhdsWithin' + apply nhdsWithin_mono _ _ hv + simp only [image_insert_eq] + apply insert_subset_insert + exact image_subset_iff.mpr st + · have : AnalyticOn 𝕜 f w := by + have : AnalyticOn 𝕜 (fun y ↦ (continuousMultilinearCurryFin0 𝕜 E F).symm (f y)) w := + ((h'p 0).mono wu).congr fun y hy ↦ (hp.zero_eq' (wu hy)).symm + have : AnalyticOn 𝕜 (fun y ↦ (continuousMultilinearCurryFin0 𝕜 E F) + ((continuousMultilinearCurryFin0 𝕜 E F).symm (f y))) w := + AnalyticOnNhd.comp_analyticOn (LinearIsometryEquiv.analyticOnNhd _ _ ) this + (mapsTo_univ _ _) + simpa using this + exact analyticOn_taylorComp h'q (fun n ↦ (h'p n).mono wu) this wv + | (n : ℕ∞) => + intro m hm + rcases hf m hm with ⟨u, hu, p, hp⟩ + rcases hg m hm with ⟨v, hv, q, hq⟩ + let w := insert x s ∩ (u ∩ f ⁻¹' v) + have wv : w ⊆ f ⁻¹' v := fun y hy => hy.2.2 + have wu : w ⊆ u := fun y hy => hy.2.1 + refine ⟨w, ?_, fun y ↦ (q (f y)).taylorComp (p y), hq.comp (hp.mono wu) wv⟩ + apply inter_mem self_mem_nhdsWithin (inter_mem hu ?_) + apply (continuousWithinAt_insert_self.2 hf.continuousWithinAt).preimage_mem_nhdsWithin' + apply nhdsWithin_mono _ _ hv + simp only [image_insert_eq] + apply insert_subset_insert + exact image_subset_iff.mpr st /-- The composition of `C^n` functions on domains is `C^n`. -/ theorem ContDiffOn.comp {s : Set E} {t : Set F} {g : F → G} {f : E → F} (hg : ContDiffOn 𝕜 n g t) @@ -889,7 +936,7 @@ theorem iteratedFDerivWithin_clm_apply_const_apply induction i generalizing x with | zero => simp | succ i ih => - replace hi : i < n := lt_of_lt_of_le (by norm_cast; simp) hi + replace hi : (i : WithTop ℕ∞) < n := lt_of_lt_of_le (by norm_cast; simp) hi have h_deriv_apply : DifferentiableOn 𝕜 (iteratedFDerivWithin 𝕜 i (fun y ↦ (c y) u) s) s := (hc.clm_apply contDiffOn_const).differentiableOn_iteratedFDerivWithin hi hs have h_deriv : DifferentiableOn 𝕜 (iteratedFDerivWithin 𝕜 i c s) s := @@ -931,9 +978,9 @@ theorem contDiff_prodAssoc_symm : ContDiff 𝕜 ⊤ <| (Equiv.prodAssoc E F G).s /-! ### Bundled derivatives are smooth -/ -/-- One direction of `contDiffWithinAt_succ_iff_hasFDerivWithinAt`, but where all derivatives -taken within the same set. Version for partial derivatives / functions with parameters. `f x` is a -`C^n+1` family of functions and `g x` is a `C^n` family of points, then the derivative of `f x` at +/-- One direction of `contDiffWithinAt_succ_iff_hasFDerivWithinAt`, but where all derivatives are +taken within the same set. Version for partial derivatives / functions with parameters. If `f x` is +a `C^n+1` family of functions and `g x` is a `C^n` family of points, then the derivative of `f x` at `g x` depends in a `C^n` way on `x`. We give a general version of this fact relative to sets which may not have unique derivatives, in the following form. If `f : E × F → G` is `C^n+1` at `(x₀, g(x₀))` in `(s ∪ {x₀}) × t ⊆ E × F` and `g : E → F` is `C^n` at `x₀` within some set `s ⊆ E`, @@ -942,7 +989,7 @@ sufficiently close to `x₀` within `s ∪ {x₀}` the function `y ↦ f x y` ha within `t ⊆ F`. For convenience, we return an explicit set of `x`'s where this holds that is a subset of `s ∪ {x₀}`. We need one additional condition, namely that `t` is a neighborhood of `g(x₀)` within `g '' s`. -/ -theorem ContDiffWithinAt.hasFDerivWithinAt_nhds {f : E → F → G} {g : E → F} {t : Set F} {n : ℕ} +theorem ContDiffWithinAt.hasFDerivWithinAt_nhds {f : E → F → G} {g : E → F} {t : Set F} (hn : n ≠ ∞) {x₀ : E} (hf : ContDiffWithinAt 𝕜 (n + 1) (uncurry f) (insert x₀ s ×ˢ t) (x₀, g x₀)) (hg : ContDiffWithinAt 𝕜 n g s x₀) (hgt : t ∈ 𝓝[g '' s] g x₀) : ∃ v ∈ 𝓝[insert x₀ s] x₀, v ⊆ insert x₀ s ∧ ∃ f' : E → F →L[𝕜] G, @@ -952,7 +999,8 @@ theorem ContDiffWithinAt.hasFDerivWithinAt_nhds {f : E → F → G} {g : E → F refine nhdsWithin_mono _ ?_ (nhdsWithin_prod self_mem_nhdsWithin hgt) simp_rw [image_subset_iff, mk_preimage_prod, preimage_id', subset_inter_iff, subset_insert, true_and, subset_preimage_image] - obtain ⟨v, hv, hvs, f', hvf', hf'⟩ := contDiffWithinAt_succ_iff_hasFDerivWithinAt'.mp hf + obtain ⟨v, hv, hvs, f_an, f', hvf', hf'⟩ := + (contDiffWithinAt_succ_iff_hasFDerivWithinAt' hn).mp hf refine ⟨(fun z => (z, g z)) ⁻¹' v ∩ insert x₀ s, ?_, inter_subset_right, fun z => (f' (z, g z)).comp (ContinuousLinearMap.inr 𝕜 E F), ?_, ?_⟩ @@ -983,21 +1031,24 @@ theorem ContDiffWithinAt.fderivWithin'' {f : E → F → G} {g : E → F} {t : S (ht : ∀ᶠ x in 𝓝[insert x₀ s] x₀, UniqueDiffWithinAt 𝕜 t (g x)) (hmn : m + 1 ≤ n) (hgt : t ∈ 𝓝[g '' s] g x₀) : ContDiffWithinAt 𝕜 m (fun x => fderivWithin 𝕜 (f x) t (g x)) s x₀ := by - have : ∀ k : ℕ, (k : ℕ∞) ≤ m → - ContDiffWithinAt 𝕜 k (fun x => fderivWithin 𝕜 (f x) t (g x)) s x₀ := fun k hkm ↦ by + have : ∀ k : ℕ, k ≤ m → ContDiffWithinAt 𝕜 k (fun x => fderivWithin 𝕜 (f x) t (g x)) s x₀ := by + intro k hkm obtain ⟨v, hv, -, f', hvf', hf'⟩ := - (hf.of_le <| (add_le_add_right hkm 1).trans hmn).hasFDerivWithinAt_nhds (hg.of_le hkm) hgt + (hf.of_le <| (add_le_add_right hkm 1).trans hmn).hasFDerivWithinAt_nhds (by simp) + (hg.of_le hkm) hgt refine hf'.congr_of_eventuallyEq_insert ?_ filter_upwards [hv, ht] exact fun y hy h2y => (hvf' y hy).fderivWithin h2y match m with | ω => - intro k hk - apply this k hk - exact le_rfl + obtain rfl : n = ω := by simpa using hmn + obtain ⟨v, hv, -, f', hvf', hf'⟩ := hf.hasFDerivWithinAt_nhds (by simp) hg hgt + refine hf'.congr_of_eventuallyEq_insert ?_ + filter_upwards [hv, ht] + exact fun y hy h2y => (hvf' y hy).fderivWithin h2y | ∞ => - rw [contDiffWithinAt_top] - exact fun m => this m (mod_cast le_top) + rw [contDiffWithinAt_infty] + exact fun k ↦ this k (by exact_mod_cast le_top) | (m : ℕ) => exact this _ le_rfl /-- A special case of `ContDiffWithinAt.fderivWithin''` where we require that `s ⊆ g⁻¹(t)`. -/ @@ -1038,7 +1089,7 @@ theorem ContDiffWithinAt.fderivWithin_right (hf : ContDiffWithinAt 𝕜 n f s x /-- `x ↦ fderivWithin 𝕜 f s x (k x)` is smooth at `x₀` within `s`. -/ theorem ContDiffWithinAt.fderivWithin_right_apply - {f : F → G} {k : F → F} {s : Set F} {n : ℕ∞} {x₀ : F} + {f : F → G} {k : F → F} {s : Set F} {x₀ : F} (hf : ContDiffWithinAt 𝕜 n f s x₀) (hk : ContDiffWithinAt 𝕜 m k s x₀) (hs : UniqueDiffOn 𝕜 s) (hmn : m + 1 ≤ n) (hx₀s : x₀ ∈ s) : ContDiffWithinAt 𝕜 m (fun x => fderivWithin 𝕜 f s x (k x)) s x₀ := @@ -1141,7 +1192,6 @@ theorem hasFTaylorSeriesUpToOn_pi {n : WithTop ℕ∞} : (fun x m => ContinuousMultilinearMap.pi fun i => p' i x m) s ↔ ∀ i, HasFTaylorSeriesUpToOn n (φ i) (p' i) s := by set pr := @ContinuousLinearMap.proj 𝕜 _ ι F' _ _ _ - letI : ∀ (m : ℕ) (i : ι), NormedSpace 𝕜 (E[×m]→L[𝕜] F' i) := fun m i => inferInstance set L : ∀ m : ℕ, (∀ i, E[×m]→L[𝕜] F' i) ≃ₗᵢ[𝕜] E[×m]→L[𝕜] ∀ i, F' i := fun m => ContinuousMultilinearMap.piₗᵢ _ _ refine ⟨fun h i => ?_, fun h => ⟨fun x hx => ?_, ?_, ?_⟩⟩ @@ -1165,10 +1215,22 @@ theorem hasFTaylorSeriesUpToOn_pi' {n : WithTop ℕ∞} : theorem contDiffWithinAt_pi : ContDiffWithinAt 𝕜 n Φ s x ↔ ∀ i, ContDiffWithinAt 𝕜 n (fun x => Φ x i) s x := by set pr := @ContinuousLinearMap.proj 𝕜 _ ι F' _ _ _ - refine ⟨fun h i => h.continuousLinearMap_comp (pr i), fun h m hm => ?_⟩ - choose u hux p hp using fun i => h i m hm - exact ⟨⋂ i, u i, Filter.iInter_mem.2 hux, _, - hasFTaylorSeriesUpToOn_pi.2 fun i => (hp i).mono <| iInter_subset _ _⟩ + refine ⟨fun h i => h.continuousLinearMap_comp (pr i), fun h ↦ ?_⟩ + match n with + | ω => + choose u hux p hp h'p using h + refine ⟨⋂ i, u i, Filter.iInter_mem.2 hux, _, + hasFTaylorSeriesUpToOn_pi.2 fun i => (hp i).mono <| iInter_subset _ _, fun m ↦ ?_⟩ + set L : (∀ i, E[×m]→L[𝕜] F' i) ≃ₗᵢ[𝕜] E[×m]→L[𝕜] ∀ i, F' i := + ContinuousMultilinearMap.piₗᵢ _ _ + change AnalyticOn 𝕜 (fun x ↦ L (fun i ↦ p i x m)) (⋂ i, u i) + apply (L.analyticOnNhd univ).comp_analyticOn ?_ (mapsTo_univ _ _) + exact AnalyticOn.pi (fun i ↦ (h'p i m).mono (iInter_subset _ _)) + | (n : ℕ∞) => + intro m hm + choose u hux p hp using fun i => h i m hm + exact ⟨⋂ i, u i, Filter.iInter_mem.2 hux, _, + hasFTaylorSeriesUpToOn_pi.2 fun i => (hp i).mono <| iInter_subset _ _⟩ theorem contDiffOn_pi : ContDiffOn 𝕜 n Φ s ↔ ∀ i, ContDiffOn 𝕜 n (fun x => Φ x i) s := ⟨fun h _ x hx => contDiffWithinAt_pi.1 (h x hx) _, fun h x hx => @@ -1627,7 +1689,9 @@ theorem contDiff_prod_mk_right (e₀ : E) : ContDiff 𝕜 n fun f : F => (e₀, end prodMap -/-! ### Inversion in a complete normed algebra -/ +/-! +### Inversion in a complete normed algebra (or more generally with summable geometric series) +-/ section AlgebraInverse @@ -1637,36 +1701,14 @@ variable {R : Type*} [NormedRing R] [NormedAlgebra 𝕜 R] open NormedRing ContinuousLinearMap Ring /-- In a complete normed algebra, the operation of inversion is `C^n`, for all `n`, at each -invertible element. The proof is by induction, bootstrapping using an identity expressing the -derivative of inversion as a bilinear map of inversion itself. -/ -theorem contDiffAt_ring_inverse [CompleteSpace R] (x : Rˣ) : +invertible element, as it is analytic. -/ +theorem contDiffAt_ring_inverse [HasSummableGeomSeries R] (x : Rˣ) : ContDiffAt 𝕜 n Ring.inverse (x : R) := by - suffices H : ∀ (n : ℕ∞), ContDiffAt 𝕜 n Ring.inverse (x : R) by - intro k hk - exact H ⊤ k (mod_cast le_top) - intro n - induction' n using ENat.nat_induction with n IH Itop - · intro m hm - refine ⟨{ y : R | IsUnit y }, ?_, ?_⟩ - · simpa [nhdsWithin_univ] using x.nhds - · use ftaylorSeriesWithin 𝕜 inverse univ - have : (m : WithTop ℕ∞) = 0 := by exact_mod_cast le_antisymm hm bot_le - rw [this, hasFTaylorSeriesUpToOn_zero_iff] - constructor - · rintro _ ⟨x', rfl⟩ - exact (inverse_continuousAt x').continuousWithinAt - · simp [ftaylorSeriesWithin] - · rw [show ((n.succ : ℕ∞) : WithTop ℕ∞) = n + 1 from rfl, contDiffAt_succ_iff_hasFDerivAt] - refine ⟨fun x : R => -mulLeftRight 𝕜 R (inverse x) (inverse x), ?_, ?_⟩ - · refine ⟨{ y : R | IsUnit y }, x.nhds, ?_⟩ - rintro _ ⟨y, rfl⟩ - simp_rw [inverse_unit] - exact hasFDerivAt_ring_inverse y - · convert (mulLeftRight_isBoundedBilinear 𝕜 R).contDiff.neg.comp_contDiffAt (x : R) - (IH.prod IH) - · exact contDiffAt_top.mpr Itop - -variable {𝕜' : Type*} [NormedField 𝕜'] [NormedAlgebra 𝕜 𝕜'] [CompleteSpace 𝕜'] + have := AnalyticOnNhd.contDiffOn (analyticOnNhd_inverse (𝕜 := 𝕜) (A := R)) (n := n) + Units.isOpen.uniqueDiffOn x x.isUnit + exact this.contDiffAt (Units.isOpen.mem_nhds x.isUnit) + +variable {𝕜' : Type*} [NormedField 𝕜'] [NormedAlgebra 𝕜 𝕜'] theorem contDiffAt_inv {x : 𝕜'} (hx : x ≠ 0) {n} : ContDiffAt 𝕜 n Inv.inv x := by simpa only [Ring.inverse_eq_inv'] using contDiffAt_ring_inverse 𝕜 (Units.mk0 x hx) @@ -1676,10 +1718,7 @@ theorem contDiffOn_inv {n} : ContDiffOn 𝕜 n (Inv.inv : 𝕜' → 𝕜') {0} variable {𝕜} --- TODO: the next few lemmas don't need `𝕜` or `𝕜'` to be complete --- A good way to show this is to generalize `contDiffAt_ring_inverse` to the setting --- of a function `f` such that `∀ᶠ x in 𝓝 a, x * f x = 1`. -theorem ContDiffWithinAt.inv {f : E → 𝕜'} (hf : ContDiffWithinAt 𝕜 n f s x) (hx : f x ≠ 0) : +theorem ContDiffWithinAt.inv {f : E → 𝕜'} {n} (hf : ContDiffWithinAt 𝕜 n f s x) (hx : f x ≠ 0) : ContDiffWithinAt 𝕜 n (fun x => (f x)⁻¹) s x := (contDiffAt_inv 𝕜 hx).comp_contDiffWithinAt x hf @@ -1695,20 +1734,20 @@ theorem ContDiff.inv {f : E → 𝕜'} (hf : ContDiff 𝕜 n f) (h : ∀ x, f x rw [contDiff_iff_contDiffAt]; exact fun x => hf.contDiffAt.inv (h x) -- TODO: generalize to `f g : E → 𝕜'` -theorem ContDiffWithinAt.div [CompleteSpace 𝕜] {f g : E → 𝕜} (hf : ContDiffWithinAt 𝕜 n f s x) +theorem ContDiffWithinAt.div {f g : E → 𝕜} {n} (hf : ContDiffWithinAt 𝕜 n f s x) (hg : ContDiffWithinAt 𝕜 n g s x) (hx : g x ≠ 0) : ContDiffWithinAt 𝕜 n (fun x => f x / g x) s x := by simpa only [div_eq_mul_inv] using hf.mul (hg.inv hx) -theorem ContDiffOn.div [CompleteSpace 𝕜] {f g : E → 𝕜} (hf : ContDiffOn 𝕜 n f s) +theorem ContDiffOn.div {f g : E → 𝕜} {n} (hf : ContDiffOn 𝕜 n f s) (hg : ContDiffOn 𝕜 n g s) (h₀ : ∀ x ∈ s, g x ≠ 0) : ContDiffOn 𝕜 n (f / g) s := fun x hx => (hf x hx).div (hg x hx) (h₀ x hx) -nonrec theorem ContDiffAt.div [CompleteSpace 𝕜] {f g : E → 𝕜} (hf : ContDiffAt 𝕜 n f x) +nonrec theorem ContDiffAt.div {f g : E → 𝕜} {n} (hf : ContDiffAt 𝕜 n f x) (hg : ContDiffAt 𝕜 n g x) (hx : g x ≠ 0) : ContDiffAt 𝕜 n (fun x => f x / g x) x := hf.div hg hx -theorem ContDiff.div [CompleteSpace 𝕜] {f g : E → 𝕜} (hf : ContDiff 𝕜 n f) (hg : ContDiff 𝕜 n g) +theorem ContDiff.div {f g : E → 𝕜} {n} (hf : ContDiff 𝕜 n f) (hg : ContDiff 𝕜 n g) (h0 : ∀ x, g x ≠ 0) : ContDiff 𝕜 n fun x => f x / g x := by simp only [contDiff_iff_contDiffAt] at * exact fun x => (hf x).div (hg x) (h0 x) @@ -1746,53 +1785,6 @@ section FunctionInverse open ContinuousLinearMap -private theorem PartialHomeomorph.contDiffAt_symm_aux {n : ℕ∞} - [CompleteSpace E] (f : PartialHomeomorph E F) - {f₀' : E ≃L[𝕜] F} {a : F} (ha : a ∈ f.target) - (hf₀' : HasFDerivAt f (f₀' : E →L[𝕜] F) (f.symm a)) (hf : ContDiffAt 𝕜 n f (f.symm a)) : - ContDiffAt 𝕜 n f.symm a := by - -- We prove this by induction on `n` - induction' n using ENat.nat_induction with n IH Itop - · apply contDiffAt_zero.2 - exact ⟨f.target, IsOpen.mem_nhds f.open_target ha, f.continuousOn_invFun⟩ - · obtain ⟨f', ⟨u, hu, hff'⟩, hf'⟩ := contDiffAt_succ_iff_hasFDerivAt.mp hf - rw [show ((n.succ : ℕ∞) : WithTop ℕ∞) = n + 1 from rfl, contDiffAt_succ_iff_hasFDerivAt] - -- For showing `n.succ` times continuous differentiability (the main inductive step), it - -- suffices to produce the derivative and show that it is `n` times continuously differentiable - have eq_f₀' : f' (f.symm a) = f₀' := (hff' (f.symm a) (mem_of_mem_nhds hu)).unique hf₀' - -- This follows by a bootstrapping formula expressing the derivative as a function of `f` itself - refine ⟨inverse ∘ f' ∘ f.symm, ?_, ?_⟩ - · -- We first check that the derivative of `f` is that formula - have h_nhds : { y : E | ∃ e : E ≃L[𝕜] F, ↑e = f' y } ∈ 𝓝 (f.symm a) := by - have hf₀' := f₀'.nhds - rw [← eq_f₀'] at hf₀' - exact hf'.continuousAt.preimage_mem_nhds hf₀' - obtain ⟨t, htu, ht, htf⟩ := mem_nhds_iff.mp (Filter.inter_mem hu h_nhds) - use f.target ∩ f.symm ⁻¹' t - refine ⟨IsOpen.mem_nhds ?_ ?_, ?_⟩ - · exact f.isOpen_inter_preimage_symm ht - · exact mem_inter ha (mem_preimage.mpr htf) - intro x hx - obtain ⟨hxu, e, he⟩ := htu hx.2 - have h_deriv : HasFDerivAt f (e : E →L[𝕜] F) (f.symm x) := by - rw [he] - exact hff' (f.symm x) hxu - convert f.hasFDerivAt_symm hx.1 h_deriv - simp [← he] - · -- Then we check that the formula, being a composition of `ContDiff` pieces, is - -- itself `ContDiff` - have h_deriv₁ : ContDiffAt 𝕜 n inverse (f' (f.symm a)) := by - rw [eq_f₀'] - exact contDiffAt_map_inverse _ - have h_deriv₂ : ContDiffAt 𝕜 n f.symm a := by - refine IH (hf.of_le ?_) - norm_cast - exact Nat.le_succ n - exact (h_deriv₁.comp _ hf').comp _ h_deriv₂ - · refine contDiffAt_top.mpr ?_ - intro n - exact Itop n (contDiffAt_top.mp hf n) - /-- If `f` is a local homeomorphism and the point `a` is in its target, and if `f` is `n` times continuously differentiable at `f.symm a`, and if the derivative at `f.symm a` is a continuous linear equivalence, @@ -1806,10 +1798,52 @@ theorem PartialHomeomorph.contDiffAt_symm [CompleteSpace E] (f : PartialHomeomor ContDiffAt 𝕜 n f.symm a := by match n with | ω => - intro k hk - exact f.contDiffAt_symm_aux ha hf₀' (hf.of_le (m := k) le_top) k le_rfl + apply AnalyticAt.contDiffAt + exact f.analyticAt_symm ha hf.analyticAt hf₀'.fderiv | (n : ℕ∞) => - exact f.contDiffAt_symm_aux ha hf₀' hf + -- We prove this by induction on `n` + induction' n using ENat.nat_induction with n IH Itop + · apply contDiffAt_zero.2 + exact ⟨f.target, IsOpen.mem_nhds f.open_target ha, f.continuousOn_invFun⟩ + · obtain ⟨f', ⟨u, hu, hff'⟩, hf'⟩ := contDiffAt_succ_iff_hasFDerivAt.mp hf + apply contDiffAt_succ_iff_hasFDerivAt.2 + -- For showing `n.succ` times continuous differentiability (the main inductive step), it + -- suffices to produce the derivative and show that it is `n` times continuously + -- differentiable + have eq_f₀' : f' (f.symm a) = f₀' := (hff' (f.symm a) (mem_of_mem_nhds hu)).unique hf₀' + -- This follows by a bootstrapping formula expressing the derivative as a + -- function of `f` itself + refine ⟨inverse ∘ f' ∘ f.symm, ?_, ?_⟩ + · -- We first check that the derivative of `f` is that formula + have h_nhds : { y : E | ∃ e : E ≃L[𝕜] F, ↑e = f' y } ∈ 𝓝 (f.symm a) := by + have hf₀' := f₀'.nhds + rw [← eq_f₀'] at hf₀' + exact hf'.continuousAt.preimage_mem_nhds hf₀' + obtain ⟨t, htu, ht, htf⟩ := mem_nhds_iff.mp (Filter.inter_mem hu h_nhds) + use f.target ∩ f.symm ⁻¹' t + refine ⟨IsOpen.mem_nhds ?_ ?_, ?_⟩ + · exact f.isOpen_inter_preimage_symm ht + · exact mem_inter ha (mem_preimage.mpr htf) + intro x hx + obtain ⟨hxu, e, he⟩ := htu hx.2 + have h_deriv : HasFDerivAt f (e : E →L[𝕜] F) (f.symm x) := by + rw [he] + exact hff' (f.symm x) hxu + convert f.hasFDerivAt_symm hx.1 h_deriv + simp [← he] + · -- Then we check that the formula, being a composition of `ContDiff` pieces, is + -- itself `ContDiff` + have h_deriv₁ : ContDiffAt 𝕜 n inverse (f' (f.symm a)) := by + rw [eq_f₀'] + exact contDiffAt_map_inverse _ + have h_deriv₂ : ContDiffAt 𝕜 n f.symm a := by + refine IH (hf.of_le ?_) + norm_cast + exact Nat.le_succ n + exact (h_deriv₁.comp _ hf').comp _ h_deriv₂ + · refine contDiffAt_infty.mpr ?_ + intro n + exact Itop n (contDiffAt_infty.mp hf n) /-- If `f` is an `n` times continuously differentiable homeomorphism, and if the derivative of `f` at each point is a continuous linear equivalence, @@ -1854,22 +1888,23 @@ variable (𝕜) that consist of points `x ∈ f.source`, `y = f x ∈ f.target` such that `f` is `C^n` at `x` and `f.symm` is `C^n` at `y`. -Note that `n` is a natural number, not `∞`, +Note that `n` is a natural number or `ω`, but not `∞`, because the set of points of `C^∞`-smoothness of `f` is not guaranteed to be open. -/ @[simps! apply symm_apply source target] -def restrContDiff (f : PartialHomeomorph E F) (n : ℕ) : PartialHomeomorph E F := +def restrContDiff (f : PartialHomeomorph E F) (n : WithTop ℕ∞) (hn : n ≠ ∞) : + PartialHomeomorph E F := haveI H : f.IsImage {x | ContDiffAt 𝕜 n f x ∧ ContDiffAt 𝕜 n f.symm (f x)} {y | ContDiffAt 𝕜 n f.symm y ∧ ContDiffAt 𝕜 n f (f.symm y)} := fun x hx ↦ by simp [hx, and_comm] H.restr <| isOpen_iff_mem_nhds.2 fun _ ⟨hxs, hxf, hxf'⟩ ↦ - inter_mem (f.open_source.mem_nhds hxs) <| hxf.eventually.and <| - f.continuousAt hxs hxf'.eventually + inter_mem (f.open_source.mem_nhds hxs) <| (hxf.eventually hn).and <| + f.continuousAt hxs (hxf'.eventually hn) -lemma contDiffOn_restrContDiff_source (f : PartialHomeomorph E F) (n : ℕ) : - ContDiffOn 𝕜 n f (f.restrContDiff 𝕜 n).source := fun _x hx ↦ hx.2.1.contDiffWithinAt +lemma contDiffOn_restrContDiff_source (f : PartialHomeomorph E F) {n : WithTop ℕ∞} (hn : n ≠ ∞) : + ContDiffOn 𝕜 n f (f.restrContDiff 𝕜 n hn).source := fun _x hx ↦ hx.2.1.contDiffWithinAt -lemma contDiffOn_restrContDiff_target (f : PartialHomeomorph E F) (n : ℕ) : - ContDiffOn 𝕜 n f.symm (f.restrContDiff 𝕜 n).target := fun _x hx ↦ hx.2.1.contDiffWithinAt +lemma contDiffOn_restrContDiff_target (f : PartialHomeomorph E F) {n : WithTop ℕ∞} (hn : n ≠ ∞) : + ContDiffOn 𝕜 n f.symm (f.restrContDiff 𝕜 n hn).target := fun _x hx ↦ hx.2.1.contDiffWithinAt end PartialHomeomorph @@ -1892,19 +1927,22 @@ open ContinuousLinearMap (smulRight) /-- A function is `C^(n + 1)` on a domain with unique derivatives if and only if it is differentiable there, and its derivative (formulated with `derivWithin`) is `C^n`. -/ -theorem contDiffOn_succ_iff_derivWithin {n : ℕ} (hs : UniqueDiffOn 𝕜 s₂) : +theorem contDiffOn_succ_iff_derivWithin (hs : UniqueDiffOn 𝕜 s₂) : ContDiffOn 𝕜 (n + 1) f₂ s₂ ↔ - DifferentiableOn 𝕜 f₂ s₂ ∧ ContDiffOn 𝕜 n (derivWithin f₂ s₂) s₂ := by + DifferentiableOn 𝕜 f₂ s₂ ∧ (n = ω → AnalyticOn 𝕜 f₂ s₂) ∧ + ContDiffOn 𝕜 n (derivWithin f₂ s₂) s₂ := by rw [contDiffOn_succ_iff_fderivWithin hs, and_congr_right_iff] intro _ constructor - · intro h + · rintro ⟨h', h⟩ + refine ⟨h', ?_⟩ have : derivWithin f₂ s₂ = (fun u : 𝕜 →L[𝕜] F => u 1) ∘ fderivWithin 𝕜 f₂ s₂ := by ext x; rfl simp_rw [this] apply ContDiff.comp_contDiffOn _ h exact (isBoundedBilinearMap_apply.isBoundedLinearMap_left _).contDiff - · intro h + · rintro ⟨h', h⟩ + refine ⟨h', ?_⟩ have : fderivWithin 𝕜 f₂ s₂ = smulRight (1 : 𝕜 →L[𝕜] 𝕜) ∘ derivWithin f₂ s₂ := by ext x; simp [derivWithin] simp only [this] @@ -1912,90 +1950,82 @@ theorem contDiffOn_succ_iff_derivWithin {n : ℕ} (hs : UniqueDiffOn 𝕜 s₂) have : IsBoundedBilinearMap 𝕜 fun _ : (𝕜 →L[𝕜] 𝕜) × F => _ := isBoundedBilinearMap_smulRight exact (this.isBoundedLinearMap_right _).contDiff +theorem contDiffOn_infty_iff_derivWithin (hs : UniqueDiffOn 𝕜 s₂) : + ContDiffOn 𝕜 ∞ f₂ s₂ ↔ DifferentiableOn 𝕜 f₂ s₂ ∧ ContDiffOn 𝕜 ∞ (derivWithin f₂ s₂) s₂ := by + rw [show ∞ = ∞ + 1 by rfl, contDiffOn_succ_iff_derivWithin hs] + simp + +@[deprecated (since := "2024-11-27")] +alias contDiffOn_top_iff_derivWithin := contDiffOn_infty_iff_derivWithin + /-- A function is `C^(n + 1)` on an open domain if and only if it is differentiable there, and its derivative (formulated with `deriv`) is `C^n`. -/ -theorem contDiffOn_succ_iff_deriv_of_isOpen {n : ℕ} (hs : IsOpen s₂) : - ContDiffOn 𝕜 (n + 1) f₂ s₂ ↔ DifferentiableOn 𝕜 f₂ s₂ ∧ ContDiffOn 𝕜 n (deriv f₂) s₂ := by +theorem contDiffOn_succ_iff_deriv_of_isOpen (hs : IsOpen s₂) : + ContDiffOn 𝕜 (n + 1) f₂ s₂ ↔ + DifferentiableOn 𝕜 f₂ s₂ ∧ (n = ω → AnalyticOn 𝕜 f₂ s₂) ∧ + ContDiffOn 𝕜 n (deriv f₂) s₂ := by rw [contDiffOn_succ_iff_derivWithin hs.uniqueDiffOn] - exact Iff.rfl.and (contDiffOn_congr fun _ => derivWithin_of_isOpen hs) + exact Iff.rfl.and (Iff.rfl.and (contDiffOn_congr fun _ => derivWithin_of_isOpen hs)) -/-- A function is `C^∞` on a domain with unique derivatives if and only if it is differentiable -there, and its derivative (formulated with `derivWithin`) is `C^∞`. -/ -theorem contDiffOn_top_iff_derivWithin (hs : UniqueDiffOn 𝕜 s₂) : - ContDiffOn 𝕜 ∞ f₂ s₂ ↔ DifferentiableOn 𝕜 f₂ s₂ ∧ ContDiffOn 𝕜 ∞ (derivWithin f₂ s₂) s₂ := by - constructor - · intro h - refine ⟨h.differentiableOn (mod_cast le_top), ?_⟩ - refine contDiffOn_top.2 fun n => ((contDiffOn_succ_iff_derivWithin hs).1 ?_).2 - exact h.of_le (mod_cast le_top) - · intro h - refine contDiffOn_top.2 fun n => ?_ - have A : (n : ℕ∞) ≤ ∞ := mod_cast le_top - apply ((contDiffOn_succ_iff_derivWithin hs).2 ⟨h.1, h.2.of_le A⟩).of_le - exact_mod_cast (Nat.le_succ n) - -/-- A function is `C^∞` on an open domain if and only if it is differentiable -there, and its derivative (formulated with `deriv`) is `C^∞`. -/ -theorem contDiffOn_top_iff_deriv_of_isOpen (hs : IsOpen s₂) : +theorem contDiffOn_infty_iff_deriv_of_isOpen (hs : IsOpen s₂) : ContDiffOn 𝕜 ∞ f₂ s₂ ↔ DifferentiableOn 𝕜 f₂ s₂ ∧ ContDiffOn 𝕜 ∞ (deriv f₂) s₂ := by - rw [contDiffOn_top_iff_derivWithin hs.uniqueDiffOn] - exact Iff.rfl.and <| contDiffOn_congr fun _ => derivWithin_of_isOpen hs + rw [show ∞ = ∞ + 1 by rfl, contDiffOn_succ_iff_deriv_of_isOpen hs] + simp + +@[deprecated (since := "2024-11-27")] +alias contDiffOn_top_iff_deriv_of_isOpen := contDiffOn_infty_iff_deriv_of_isOpen protected theorem ContDiffOn.derivWithin (hf : ContDiffOn 𝕜 n f₂ s₂) (hs : UniqueDiffOn 𝕜 s₂) - (hmn : m + 1 ≤ n) : ContDiffOn 𝕜 m (derivWithin f₂ s₂) s₂ := by - rcases le_or_lt ∞ n with hn | hn - · have : ContDiffOn 𝕜 ∞ (derivWithin f₂ s₂) s₂ := - ((contDiffOn_top_iff_derivWithin hs).1 (hf.of_le hn)).2 - intro x hx k hk - exact this x hx k (mod_cast le_top) - · match m with - | ω => simpa using hmn.trans_lt hn - | ∞ => simpa using hmn.trans_lt hn - | (m : ℕ) => - change (m.succ : ℕ∞) ≤ n at hmn - exact ((contDiffOn_succ_iff_derivWithin hs).1 (hf.of_le hmn)).2 + (hmn : m + 1 ≤ n) : ContDiffOn 𝕜 m (derivWithin f₂ s₂) s₂ := + ((contDiffOn_succ_iff_derivWithin hs).1 (hf.of_le hmn)).2.2 theorem ContDiffOn.deriv_of_isOpen (hf : ContDiffOn 𝕜 n f₂ s₂) (hs : IsOpen s₂) (hmn : m + 1 ≤ n) : ContDiffOn 𝕜 m (deriv f₂) s₂ := (hf.derivWithin hs.uniqueDiffOn hmn).congr fun _ hx => (derivWithin_of_isOpen hs hx).symm theorem ContDiffOn.continuousOn_derivWithin (h : ContDiffOn 𝕜 n f₂ s₂) (hs : UniqueDiffOn 𝕜 s₂) - (hn : 1 ≤ n) : ContinuousOn (derivWithin f₂ s₂) s₂ := - ((contDiffOn_succ_iff_derivWithin hs).1 (h.of_le hn)).2.continuousOn + (hn : 1 ≤ n) : ContinuousOn (derivWithin f₂ s₂) s₂ := by + rw [show (1 : WithTop ℕ∞) = 0 + 1 from rfl] at hn + exact ((contDiffOn_succ_iff_derivWithin hs).1 (h.of_le hn)).2.2.continuousOn theorem ContDiffOn.continuousOn_deriv_of_isOpen (h : ContDiffOn 𝕜 n f₂ s₂) (hs : IsOpen s₂) - (hn : 1 ≤ n) : ContinuousOn (deriv f₂) s₂ := - ((contDiffOn_succ_iff_deriv_of_isOpen hs).1 (h.of_le hn)).2.continuousOn + (hn : 1 ≤ n) : ContinuousOn (deriv f₂) s₂ := by + rw [show (1 : WithTop ℕ∞) = 0 + 1 from rfl] at hn + exact ((contDiffOn_succ_iff_deriv_of_isOpen hs).1 (h.of_le hn)).2.2.continuousOn /-- A function is `C^(n + 1)` if and only if it is differentiable, and its derivative (formulated in terms of `deriv`) is `C^n`. -/ -theorem contDiff_succ_iff_deriv {n : ℕ} : - ContDiff 𝕜 (n + 1) f₂ ↔ Differentiable 𝕜 f₂ ∧ ContDiff 𝕜 n (deriv f₂) := by +theorem contDiff_succ_iff_deriv : + ContDiff 𝕜 (n + 1) f₂ ↔ Differentiable 𝕜 f₂ ∧ (n = ω → AnalyticOn 𝕜 f₂ univ) ∧ + ContDiff 𝕜 n (deriv f₂) := by simp only [← contDiffOn_univ, contDiffOn_succ_iff_deriv_of_isOpen, isOpen_univ, differentiableOn_univ] -theorem contDiff_one_iff_deriv : ContDiff 𝕜 1 f₂ ↔ Differentiable 𝕜 f₂ ∧ Continuous (deriv f₂) := - contDiff_succ_iff_deriv.trans <| Iff.rfl.and contDiff_zero +theorem contDiff_one_iff_deriv : + ContDiff 𝕜 1 f₂ ↔ Differentiable 𝕜 f₂ ∧ Continuous (deriv f₂) := by + rw [show (1 : WithTop ℕ∞) = 0 + 1 from rfl, contDiff_succ_iff_deriv] + simp -/-- A function is `C^∞` if and only if it is differentiable, -and its derivative (formulated in terms of `deriv`) is `C^∞`. -/ -theorem contDiff_top_iff_deriv : +theorem contDiff_infty_iff_deriv : ContDiff 𝕜 ∞ f₂ ↔ Differentiable 𝕜 f₂ ∧ ContDiff 𝕜 ∞ (deriv f₂) := by - simp only [← contDiffOn_univ, ← differentiableOn_univ, ← derivWithin_univ] - rw [contDiffOn_top_iff_derivWithin uniqueDiffOn_univ] + rw [show (∞ : WithTop ℕ∞) = ∞ + 1 from rfl, contDiff_succ_iff_deriv] + simp + +@[deprecated (since := "2024-11-27")] alias contDiff_top_iff_deriv := contDiff_infty_iff_deriv -theorem ContDiff.continuous_deriv (h : ContDiff 𝕜 n f₂) (hn : 1 ≤ n) : Continuous (deriv f₂) := - (contDiff_succ_iff_deriv.mp (h.of_le hn)).2.continuous +theorem ContDiff.continuous_deriv (h : ContDiff 𝕜 n f₂) (hn : 1 ≤ n) : Continuous (deriv f₂) := by + rw [show (1 : WithTop ℕ∞) = 0 + 1 from rfl] at hn + exact (contDiff_succ_iff_deriv.mp (h.of_le hn)).2.2.continuous theorem ContDiff.iterate_deriv : ∀ (n : ℕ) {f₂ : 𝕜 → F}, ContDiff 𝕜 ∞ f₂ → ContDiff 𝕜 ∞ (deriv^[n] f₂) | 0, _, hf => hf - | n + 1, _, hf => ContDiff.iterate_deriv n (contDiff_top_iff_deriv.mp hf).2 + | n + 1, _, hf => ContDiff.iterate_deriv n (contDiff_infty_iff_deriv.mp hf).2 theorem ContDiff.iterate_deriv' (n : ℕ) : ∀ (k : ℕ) {f₂ : 𝕜 → F}, ContDiff 𝕜 (n + k : ℕ) f₂ → ContDiff 𝕜 n (deriv^[k] f₂) | 0, _, hf => hf - | k + 1, _, hf => ContDiff.iterate_deriv' _ k (contDiff_succ_iff_deriv.mp hf).2 + | k + 1, _, hf => ContDiff.iterate_deriv' _ k (contDiff_succ_iff_deriv.mp hf).2.2 end deriv @@ -2029,9 +2059,18 @@ theorem HasFTaylorSeriesUpToOn.restrictScalars {n : WithTop ℕ∞} cont m hm := ContinuousMultilinearMap.continuous_restrictScalars.comp_continuousOn (h.cont m hm) theorem ContDiffWithinAt.restrict_scalars (h : ContDiffWithinAt 𝕜' n f s x) : - ContDiffWithinAt 𝕜 n f s x := fun m hm ↦ by - rcases h m hm with ⟨u, u_mem, p', hp'⟩ - exact ⟨u, u_mem, _, hp'.restrictScalars _⟩ + ContDiffWithinAt 𝕜 n f s x := by + match n with + | ω => + obtain ⟨u, u_mem, p', hp', Hp'⟩ := h + refine ⟨u, u_mem, _, hp'.restrictScalars _, fun i ↦ ?_⟩ + change AnalyticOn 𝕜 (fun x ↦ ContinuousMultilinearMap.restrictScalarsLinear 𝕜 (p' x i)) u + apply AnalyticOnNhd.comp_analyticOn _ (Hp' i).restrictScalars (Set.mapsTo_univ _ _) + exact ContinuousLinearMap.analyticOnNhd _ _ + | (n : ℕ∞) => + intro m hm + rcases h m hm with ⟨u, u_mem, p', hp'⟩ + exact ⟨u, u_mem, _, hp'.restrictScalars _⟩ theorem ContDiffOn.restrict_scalars (h : ContDiffOn 𝕜' n f s) : ContDiffOn 𝕜 n f s := fun x hx => (h x hx).restrict_scalars _ diff --git a/Mathlib/Analysis/Calculus/ContDiff/CPolynomial.lean b/Mathlib/Analysis/Calculus/ContDiff/CPolynomial.lean index 3a66c7c3d570a..ebd550a6c8643 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/CPolynomial.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/CPolynomial.lean @@ -15,7 +15,7 @@ of continuous multilinear maps. open Filter Asymptotics -open scoped ENNReal ContDiff +open scoped ENNReal universe u v @@ -32,15 +32,12 @@ variable {f : E → F} {x : E} {s : Set E} theorem CPolynomialOn.contDiffOn (h : CPolynomialOn 𝕜 f s) {n : WithTop ℕ∞} : ContDiffOn 𝕜 n f s := by let t := { x | CPolynomialAt 𝕜 f x } - suffices ContDiffOn 𝕜 ω f t from (this.of_le le_top).mono h - rw [← contDiffOn_infty_iff_contDiffOn_omega] + suffices ContDiffOn 𝕜 n f t from this.mono h + suffices AnalyticOnNhd 𝕜 f t by + have t_open : IsOpen t := isOpen_cPolynomialAt 𝕜 f + exact AnalyticOnNhd.contDiffOn this t_open.uniqueDiffOn have H : CPolynomialOn 𝕜 f t := fun _x hx ↦ hx - have t_open : IsOpen t := isOpen_cPolynomialAt 𝕜 f - exact contDiffOn_of_continuousOn_differentiableOn - (fun m _ ↦ (H.iteratedFDeriv m).continuousOn.congr - fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) - (fun m _ ↦ (H.iteratedFDeriv m).analyticOnNhd.differentiableOn.congr - fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) + exact H.analyticOnNhd theorem CPolynomialAt.contDiffAt (h : CPolynomialAt 𝕜 f x) {n : WithTop ℕ∞} : ContDiffAt 𝕜 n f x := diff --git a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean index 180ea22be1a5e..58258d4bbd896 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean @@ -3,6 +3,8 @@ Copyright (c) 2019 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ +import Mathlib.Analysis.Analytic.Within +import Mathlib.Analysis.Calculus.FDeriv.Analytic import Mathlib.Analysis.Calculus.ContDiff.FTaylorSeries /-! @@ -11,7 +13,9 @@ import Mathlib.Analysis.Calculus.ContDiff.FTaylorSeries A function is `C^1` on a domain if it is differentiable there, and its derivative is continuous. By induction, it is `C^n` if it is `C^{n-1}` and its (n-1)-th derivative is `C^1` there or, equivalently, if it is `C^1` and its derivative is `C^{n-1}`. -Finally, it is `C^∞` if it is `C^n` for all n. +It is `C^∞` if it is `C^n` for all n. +Finally, it is `C^ω` if it is analytic (as well as all its derivative, which is automatic if the +space is complete). We formalize these notions with predicates `ContDiffWithinAt`, `ContDiffAt`, `ContDiffOn` and `ContDiff` saying that the function is `C^n` within a set at a point, at a point, on a set @@ -81,7 +85,9 @@ a neighborhood of `x` within `s ∪ {x}` (which appears as `insert x s` in this We use the notation `E [×n]→L[𝕜] F` for the space of continuous multilinear maps on `E^n` with values in `F`. This is the space in which the `n`-th derivative of a function from `E` to `F` lives. -In this file, we denote `⊤ : ℕ∞` with `∞`. +In this file, we denote `(⊤ : ℕ∞) : WithTop ℕ∞` with `∞`, and `⊤ : WithTop ℕ∞` with `ω`. To +avoid ambiguities with the two tops, the theorems name use either `infty` or `omega`. +These notations are scoped in `ContDiff`. ## Tags @@ -93,14 +99,6 @@ noncomputable section open scoped Classical open NNReal Topology Filter -/-- Smoothness exponent for analytic functions. Not implemented yet, `ω` smoothness is equivalent -to `∞` smoothness in current mathlib. -/ -scoped [ContDiff] notation3 "ω" => (⊤ : WithTop ℕ∞) -/-- Smoothness exponent for infinitely differentiable functions. -/ -scoped [ContDiff] notation3 "∞" => ((⊤ : ℕ∞) : WithTop ℕ∞) - -open ContDiff - /- Porting note: These lines are not required in Mathlib4. attribute [local instance 1001] @@ -109,6 +107,13 @@ attribute [local instance 1001] open Set Fin Filter Function +/-- Smoothness exponent for analytic functions. -/ +scoped [ContDiff] notation3 "ω" => (⊤ : WithTop ℕ∞) +/-- Smoothness exponent for infinitely differentiable functions. -/ +scoped [ContDiff] notation3 "∞" => ((⊤ : ℕ∞) : WithTop ℕ∞) + +open ContDiff + universe u uE uF uG uX variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] {E : Type uE} [NormedAddCommGroup E] @@ -119,51 +124,119 @@ variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] {E : Type uE} [NormedAdd /-! ### Smooth functions within a set around a point -/ -variable (𝕜) - +variable (𝕜) in /-- A function is continuously differentiable up to order `n` within a set `s` at a point `x` if it admits continuous derivatives up to order `n` in a neighborhood of `x` in `s ∪ {x}`. For `n = ∞`, we only require that this holds up to any finite order (where the neighborhood may depend on the finite order we consider). +For `n = ω`, we require the function to be analytic within `s` at `x`. The precise definition we +give (all the derivatives should be analytic) is more involved to work around issues when the space +is not complete, but it is equivalent when the space is complete. For instance, a real function which is `C^m` on `(-1/m, 1/m)` for each natural `m`, but not better, is `C^∞` at `0` within `univ`. - -We take the exponent `n` in `WithTop ℕ∞` to allow for an extension to analytic functions in the -future, but currently the notion is the same for `n = ∞` and `n = ω`. -/ def ContDiffWithinAt (n : WithTop ℕ∞) (f : E → F) (s : Set E) (x : E) : Prop := - ∀ m : ℕ, m ≤ n → ∃ u ∈ 𝓝[insert x s] x, - ∃ p : E → FormalMultilinearSeries 𝕜 E F, HasFTaylorSeriesUpToOn m f p u - -variable {𝕜} + match n with + | ω => ∃ u ∈ 𝓝[insert x s] x, ∃ p : E → FormalMultilinearSeries 𝕜 E F, + HasFTaylorSeriesUpToOn ω f p u ∧ ∀ i, AnalyticOn 𝕜 (fun x ↦ p x i) u + | (n : ℕ∞) => ∀ m : ℕ, m ≤ n → ∃ u ∈ 𝓝[insert x s] x, + ∃ p : E → FormalMultilinearSeries 𝕜 E F, HasFTaylorSeriesUpToOn m f p u + +lemma HasFTaylorSeriesUpToOn.analyticOn + (hf : HasFTaylorSeriesUpToOn ω f p s) (h : AnalyticOn 𝕜 (fun x ↦ p x 0) s) : + AnalyticOn 𝕜 f s := by + have : AnalyticOn 𝕜 (fun x ↦ (continuousMultilinearCurryFin0 𝕜 E F) (p x 0)) s := + (LinearIsometryEquiv.analyticOnNhd _ _ ).comp_analyticOn + h (Set.mapsTo_univ _ _) + exact this.congr (fun y hy ↦ (hf.zero_eq _ hy).symm) + +lemma ContDiffWithinAt.analyticOn (h : ContDiffWithinAt 𝕜 ω f s x) : + ∃ u ∈ 𝓝[insert x s] x, AnalyticOn 𝕜 f u := by + obtain ⟨u, hu, p, hp, h'p⟩ := h + exact ⟨u, hu, hp.analyticOn (h'p 0)⟩ + +lemma ContDiffWithinAt.analyticWithinAt (h : ContDiffWithinAt 𝕜 ω f s x) : + AnalyticWithinAt 𝕜 f s x := by + obtain ⟨u, hu, hf⟩ := h.analyticOn + have xu : x ∈ u := mem_of_mem_nhdsWithin (by simp) hu + exact (hf x xu).mono_of_mem_nhdsWithin (nhdsWithin_mono _ (subset_insert _ _) hu) + +theorem contDiffWithinAt_omega_iff_analyticWithinAt [CompleteSpace F] : + ContDiffWithinAt 𝕜 ω f s x ↔ AnalyticWithinAt 𝕜 f s x := by + refine ⟨fun h ↦ h.analyticWithinAt, fun h ↦ ?_⟩ + obtain ⟨u, hu, p, hp, h'p⟩ := h.exists_hasFTaylorSeriesUpToOn ω + exact ⟨u, hu, p, hp.of_le le_top, fun i ↦ h'p i⟩ theorem contDiffWithinAt_nat {n : ℕ} : ContDiffWithinAt 𝕜 n f s x ↔ ∃ u ∈ 𝓝[insert x s] x, ∃ p : E → FormalMultilinearSeries 𝕜 E F, HasFTaylorSeriesUpToOn n f p u := ⟨fun H => H n le_rfl, fun ⟨u, hu, p, hp⟩ _m hm => ⟨u, hu, p, hp.of_le (mod_cast hm)⟩⟩ +/-- When `n` is either a natural number or `ω`, one can characterize the property of being `C^n` +as the existence of a neighborhood on which there is a Taylor series up to order `n`, +requiring in addition that its terms are analytic in the `ω` case. -/ +lemma contDiffWithinAt_iff_of_ne_infty (hn : n ≠ ∞) : + ContDiffWithinAt 𝕜 n f s x ↔ ∃ u ∈ 𝓝[insert x s] x, + ∃ p : E → FormalMultilinearSeries 𝕜 E F, HasFTaylorSeriesUpToOn n f p u ∧ + (n = ω → ∀ i, AnalyticOn 𝕜 (fun x ↦ p x i) u) := by + match n with + | ω => simp [ContDiffWithinAt] + | ∞ => simp at hn + | (n : ℕ) => simp [contDiffWithinAt_nat] + theorem ContDiffWithinAt.of_le (h : ContDiffWithinAt 𝕜 n f s x) (hmn : m ≤ n) : - ContDiffWithinAt 𝕜 m f s x := fun k hk => h k (le_trans hk hmn) + ContDiffWithinAt 𝕜 m f s x := by + match n with + | ω => match m with + | ω => exact h + | (m : ℕ∞) => + intro k _ + obtain ⟨u, hu, p, hp, -⟩ := h + exact ⟨u, hu, p, hp.of_le le_top⟩ + | (n : ℕ∞) => match m with + | ω => simp at hmn + | (m : ℕ∞) => exact fun k hk ↦ h k (le_trans hk (mod_cast hmn)) + +/-- In a complete space, a function which is analytic within a set at a point is also `C^ω` there. +Note that the same statement for `AnalyticOn` does not require completeness, see +`AnalyticOn.contDiffOn`. -/ +theorem AnalyticWithinAt.contDiffWithinAt [CompleteSpace F] (h : AnalyticWithinAt 𝕜 f s x) : + ContDiffWithinAt 𝕜 n f s x := + (contDiffWithinAt_omega_iff_analyticWithinAt.2 h).of_le le_top theorem contDiffWithinAt_iff_forall_nat_le {n : ℕ∞} : ContDiffWithinAt 𝕜 n f s x ↔ ∀ m : ℕ, ↑m ≤ n → ContDiffWithinAt 𝕜 m f s x := - ⟨fun H _m hm => H.of_le (mod_cast hm), fun H m hm => H m (mod_cast hm) _ le_rfl⟩ + ⟨fun H _ hm => H.of_le (mod_cast hm), fun H m hm => H m hm _ le_rfl⟩ -theorem contDiffWithinAt_top : ContDiffWithinAt 𝕜 ∞ f s x ↔ ∀ n : ℕ, ContDiffWithinAt 𝕜 n f s x := +theorem contDiffWithinAt_infty : + ContDiffWithinAt 𝕜 ∞ f s x ↔ ∀ n : ℕ, ContDiffWithinAt 𝕜 n f s x := contDiffWithinAt_iff_forall_nat_le.trans <| by simp only [forall_prop_of_true, le_top] +@[deprecated (since := "2024-11-25")] alias contDiffWithinAt_top := contDiffWithinAt_infty + theorem ContDiffWithinAt.continuousWithinAt (h : ContDiffWithinAt 𝕜 n f s x) : ContinuousWithinAt f s x := by - rcases h 0 bot_le with ⟨u, hu, p, H⟩ + have := h.of_le (zero_le _) + simp only [ContDiffWithinAt, nonpos_iff_eq_zero, Nat.cast_eq_zero, + mem_pure, forall_eq, CharP.cast_eq_zero] at this + rcases this with ⟨u, hu, p, H⟩ rw [mem_nhdsWithin_insert] at hu exact (H.continuousOn.continuousWithinAt hu.1).mono_of_mem_nhdsWithin hu.2 theorem ContDiffWithinAt.congr_of_eventuallyEq (h : ContDiffWithinAt 𝕜 n f s x) - (h₁ : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : ContDiffWithinAt 𝕜 n f₁ s x := fun m hm => - let ⟨u, hu, p, H⟩ := h m hm - ⟨{ x ∈ u | f₁ x = f x }, Filter.inter_mem hu (mem_nhdsWithin_insert.2 ⟨hx, h₁⟩), p, - (H.mono (sep_subset _ _)).congr fun _ => And.right⟩ + (h₁ : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : ContDiffWithinAt 𝕜 n f₁ s x := by + match n with + | ω => + obtain ⟨u, hu, p, H, H'⟩ := h + exact ⟨{x ∈ u | f₁ x = f x}, Filter.inter_mem hu (mem_nhdsWithin_insert.2 ⟨hx, h₁⟩), p, + (H.mono (sep_subset _ _)).congr fun _ ↦ And.right, + fun i ↦ (H' i).mono (sep_subset _ _)⟩ + | (n : ℕ∞) => + intro m hm + let ⟨u, hu, p, H⟩ := h m hm + exact ⟨{ x ∈ u | f₁ x = f x }, Filter.inter_mem hu (mem_nhdsWithin_insert.2 ⟨hx, h₁⟩), p, + (H.mono (sep_subset _ _)).congr fun _ ↦ And.right⟩ theorem Filter.EventuallyEq.congr_contDiffWithinAt (h₁ : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : ContDiffWithinAt 𝕜 n f₁ s x ↔ ContDiffWithinAt 𝕜 n f s x := @@ -221,9 +294,14 @@ theorem contDiffWithinAt_congr_of_insert (h₁ : ∀ y ∈ insert x s, f₁ y = theorem ContDiffWithinAt.mono_of_mem_nhdsWithin (h : ContDiffWithinAt 𝕜 n f s x) {t : Set E} (hst : s ∈ 𝓝[t] x) : ContDiffWithinAt 𝕜 n f t x := by - intro m hm - rcases h m hm with ⟨u, hu, p, H⟩ - exact ⟨u, nhdsWithin_le_of_mem (insert_mem_nhdsWithin_insert hst) hu, p, H⟩ + match n with + | ω => + obtain ⟨u, hu, p, H, H'⟩ := h + exact ⟨u, nhdsWithin_le_of_mem (insert_mem_nhdsWithin_insert hst) hu, p, H, H'⟩ + | (n : ℕ∞) => + intro m hm + rcases h m hm with ⟨u, hu, p, H⟩ + exact ⟨u, nhdsWithin_le_of_mem (insert_mem_nhdsWithin_insert hst) hu, p, H⟩ @[deprecated (since := "2024-10-30")] alias ContDiffWithinAt.mono_of_mem := ContDiffWithinAt.mono_of_mem_nhdsWithin @@ -262,13 +340,17 @@ theorem contDiffWithinAt_inter (h : t ∈ 𝓝 x) : theorem contDiffWithinAt_insert_self : ContDiffWithinAt 𝕜 n f (insert x s) x ↔ ContDiffWithinAt 𝕜 n f s x := by - simp_rw [ContDiffWithinAt, insert_idem] + match n with + | ω => simp [ContDiffWithinAt] + | (n : ℕ∞) => simp_rw [ContDiffWithinAt, insert_idem] theorem contDiffWithinAt_insert {y : E} : ContDiffWithinAt 𝕜 n f (insert y s) x ↔ ContDiffWithinAt 𝕜 n f s x := by - rcases eq_or_ne x y with (rfl | h) + rcases eq_or_ne x y with (rfl | hx) · exact contDiffWithinAt_insert_self - simp_rw [ContDiffWithinAt, insert_comm x y, nhdsWithin_insert_of_ne h] + refine ⟨fun h ↦ h.mono (subset_insert _ _), fun h ↦ ?_⟩ + apply h.mono_of_mem_nhdsWithin + simp [nhdsWithin_insert_of_ne hx, self_mem_nhdsWithin] alias ⟨ContDiffWithinAt.of_insert, ContDiffWithinAt.insert'⟩ := contDiffWithinAt_insert @@ -284,7 +366,7 @@ theorem contDiffWithinAt_diff_singleton {y : E} : within this set at this point. -/ theorem ContDiffWithinAt.differentiableWithinAt' (h : ContDiffWithinAt 𝕜 n f s x) (hn : 1 ≤ n) : DifferentiableWithinAt 𝕜 f (insert x s) x := by - rcases h 1 hn with ⟨u, hu, p, H⟩ + rcases contDiffWithinAt_nat.1 (h.of_le hn) with ⟨u, hu, p, H⟩ rcases mem_nhdsWithin.1 hu with ⟨t, t_open, xt, tu⟩ rw [inter_comm] at tu exact (differentiableWithinAt_inter (IsOpen.mem_nhds t_open xt)).1 <| @@ -297,34 +379,42 @@ theorem ContDiffWithinAt.differentiableWithinAt (h : ContDiffWithinAt 𝕜 n f s DifferentiableWithinAt 𝕜 f s x := (h.differentiableWithinAt' hn).mono (subset_insert x s) -/-- A function is `C^(n + 1)` on a domain iff locally, it has a derivative which is `C^n`. -/ -theorem contDiffWithinAt_succ_iff_hasFDerivWithinAt {n : ℕ} : - ContDiffWithinAt 𝕜 (n + 1) f s x ↔ ∃ u ∈ 𝓝[insert x s] x, ∃ f' : E → E →L[𝕜] F, +/-- A function is `C^(n + 1)` on a domain iff locally, it has a derivative which is `C^n` +(and moreover the function is analytic when `n = ω`). -/ +theorem contDiffWithinAt_succ_iff_hasFDerivWithinAt (hn : n ≠ ∞) : + ContDiffWithinAt 𝕜 (n + 1) f s x ↔ ∃ u ∈ 𝓝[insert x s] x, (n = ω → AnalyticOn 𝕜 f u) ∧ + ∃ f' : E → E →L[𝕜] F, (∀ x ∈ u, HasFDerivWithinAt f (f' x) u x) ∧ ContDiffWithinAt 𝕜 n f' u x := by + have h'n : n + 1 ≠ ∞ := by simpa using hn constructor · intro h - rcases h n.succ le_rfl with ⟨u, hu, p, Hp⟩ - refine - ⟨u, hu, fun y => (continuousMultilinearCurryFin1 𝕜 E F) (p y 1), fun y hy => - Hp.hasFDerivWithinAt (mod_cast (Nat.le_add_left 1 n)) hy, ?_⟩ - intro m hm + rcases (contDiffWithinAt_iff_of_ne_infty h'n).1 h with ⟨u, hu, p, Hp, H'p⟩ + refine ⟨u, hu, ?_, fun y => (continuousMultilinearCurryFin1 𝕜 E F) (p y 1), + fun y hy => Hp.hasFDerivWithinAt le_add_self hy, ?_⟩ + · rintro rfl + exact Hp.analyticOn (H'p rfl 0) + apply (contDiffWithinAt_iff_of_ne_infty hn).2 refine ⟨u, ?_, fun y : E => (p y).shift, ?_⟩ · -- Porting note: without the explicit argument Lean is not sure of the type. convert @self_mem_nhdsWithin _ _ x u have : x ∈ insert x s := by simp exact insert_eq_of_mem (mem_of_mem_nhdsWithin this hu) - · rw [show ((n.succ : ℕ) : WithTop ℕ∞) = n + 1 from rfl, - hasFTaylorSeriesUpToOn_succ_iff_right] at Hp - exact Hp.2.2.of_le (mod_cast hm) - · rintro ⟨u, hu, f', f'_eq_deriv, Hf'⟩ - rw [show (n : WithTop ℕ∞) + 1 = (n + 1 : ℕ) from rfl, contDiffWithinAt_nat] - rcases Hf' n le_rfl with ⟨v, hv, p', Hp'⟩ - refine ⟨v ∩ u, ?_, fun x => (p' x).unshift (f x), ?_⟩ + · rw [hasFTaylorSeriesUpToOn_succ_iff_right] at Hp + refine ⟨Hp.2.2, ?_⟩ + rintro rfl i + change AnalyticOn 𝕜 + (fun x ↦ (continuousMultilinearCurryRightEquiv' 𝕜 i E F) (p x (i + 1))) u + apply (LinearIsometryEquiv.analyticOnNhd _ _).comp_analyticOn + ?_ (Set.mapsTo_univ _ _) + exact H'p rfl _ + · rintro ⟨u, hu, hf, f', f'_eq_deriv, Hf'⟩ + rw [contDiffWithinAt_iff_of_ne_infty h'n] + rcases (contDiffWithinAt_iff_of_ne_infty hn).1 Hf' with ⟨v, hv, p', Hp', p'_an⟩ + refine ⟨v ∩ u, ?_, fun x => (p' x).unshift (f x), ?_, ?_⟩ · apply Filter.inter_mem _ hu apply nhdsWithin_le_of_mem hu exact nhdsWithin_mono _ (subset_insert x u) hv - · rw [show ((n.succ : ℕ) : WithTop ℕ∞) = n + 1 from rfl, - hasFTaylorSeriesUpToOn_succ_iff_right] + · rw [hasFTaylorSeriesUpToOn_succ_iff_right] refine ⟨fun y _ => rfl, fun y hy => ?_, ?_⟩ · change HasFDerivWithinAt (fun z => (continuousMultilinearCurryFin0 𝕜 E F).symm (f z)) @@ -347,32 +437,48 @@ theorem contDiffWithinAt_succ_iff_hasFDerivWithinAt {n : ℕ} : change p' x k (init (@snoc k (fun _ : Fin k.succ => E) v y)) (@snoc k (fun _ : Fin k.succ => E) v y (last k)) = p' x k v y rw [snoc_last, init_snoc] + · intro h i + simp only [WithTop.add_eq_top, WithTop.one_ne_top, or_false] at h + match i with + | 0 => + simp only [FormalMultilinearSeries.unshift] + apply AnalyticOnNhd.comp_analyticOn _ ((hf h).mono inter_subset_right) + (Set.mapsTo_univ _ _) + exact LinearIsometryEquiv.analyticOnNhd _ _ + | i + 1 => + simp only [FormalMultilinearSeries.unshift, Nat.succ_eq_add_one] + apply AnalyticOnNhd.comp_analyticOn _ ((p'_an h i).mono inter_subset_left) + (Set.mapsTo_univ _ _) + exact LinearIsometryEquiv.analyticOnNhd _ _ /-- A version of `contDiffWithinAt_succ_iff_hasFDerivWithinAt` where all derivatives are taken within the same set. -/ -theorem contDiffWithinAt_succ_iff_hasFDerivWithinAt' {n : ℕ} : +theorem contDiffWithinAt_succ_iff_hasFDerivWithinAt' (hn : n ≠ ∞) : ContDiffWithinAt 𝕜 (n + 1) f s x ↔ - ∃ u ∈ 𝓝[insert x s] x, u ⊆ insert x s ∧ ∃ f' : E → E →L[𝕜] F, + ∃ u ∈ 𝓝[insert x s] x, u ⊆ insert x s ∧ (n = ω → AnalyticOn 𝕜 f u) ∧ + ∃ f' : E → E →L[𝕜] F, (∀ x ∈ u, HasFDerivWithinAt f (f' x) s x) ∧ ContDiffWithinAt 𝕜 n f' s x := by refine ⟨fun hf => ?_, ?_⟩ - · obtain ⟨u, hu, f', huf', hf'⟩ := contDiffWithinAt_succ_iff_hasFDerivWithinAt.mp hf + · obtain ⟨u, hu, f_an, f', huf', hf'⟩ := (contDiffWithinAt_succ_iff_hasFDerivWithinAt hn).mp hf obtain ⟨w, hw, hxw, hwu⟩ := mem_nhdsWithin.mp hu rw [inter_comm] at hwu - refine ⟨insert x s ∩ w, inter_mem_nhdsWithin _ (hw.mem_nhds hxw), inter_subset_left, f', + refine ⟨insert x s ∩ w, inter_mem_nhdsWithin _ (hw.mem_nhds hxw), inter_subset_left, ?_, f', fun y hy => ?_, ?_⟩ + · intro h + apply (f_an h).mono hwu · refine ((huf' y <| hwu hy).mono hwu).mono_of_mem_nhdsWithin ?_ refine mem_of_superset ?_ (inter_subset_inter_left _ (subset_insert _ _)) exact inter_mem_nhdsWithin _ (hw.mem_nhds hy.2) · exact hf'.mono_of_mem_nhdsWithin (nhdsWithin_mono _ (subset_insert _ _) hu) - · rw [← contDiffWithinAt_insert, contDiffWithinAt_succ_iff_hasFDerivWithinAt, + · rw [← contDiffWithinAt_insert, contDiffWithinAt_succ_iff_hasFDerivWithinAt hn, insert_eq_of_mem (mem_insert _ _)] - rintro ⟨u, hu, hus, f', huf', hf'⟩ - exact ⟨u, hu, f', fun y hy => (huf' y hy).insert'.mono hus, hf'.insert.mono hus⟩ + rintro ⟨u, hu, hus, f_an, f', huf', hf'⟩ + exact ⟨u, hu, f_an, f', fun y hy => (huf' y hy).insert'.mono hus, hf'.insert.mono hus⟩ -/-! ### Smooth functions within a set -/ -variable (𝕜) +/-! ### Smooth functions within a set -/ +variable (𝕜) in /-- A function is continuously differentiable up to `n` on `s` if, for any point `x` in `s`, it admits continuous derivatives up to order `n` on a neighborhood of `x` in `s`. @@ -382,9 +488,7 @@ depend on the finite order we consider). def ContDiffOn (n : WithTop ℕ∞) (f : E → F) (s : Set E) : Prop := ∀ x ∈ s, ContDiffWithinAt 𝕜 n f s x -variable {𝕜} - -theorem HasFTaylorSeriesUpToOn.contDiffOn {f' : E → FormalMultilinearSeries 𝕜 E F} +theorem HasFTaylorSeriesUpToOn.contDiffOn {n : ℕ∞} {f' : E → FormalMultilinearSeries 𝕜 E F} (hf : HasFTaylorSeriesUpToOn n f f' s) : ContDiffOn 𝕜 n f s := by intro x hx m hm use s @@ -395,59 +499,80 @@ theorem ContDiffOn.contDiffWithinAt (h : ContDiffOn 𝕜 n f s) (hx : x ∈ s) : ContDiffWithinAt 𝕜 n f s x := h x hx -theorem ContDiffWithinAt.contDiffOn' {m : ℕ} (hm : m ≤ n) +theorem ContDiffOn.of_le (h : ContDiffOn 𝕜 n f s) (hmn : m ≤ n) : ContDiffOn 𝕜 m f s := fun x hx => + (h x hx).of_le hmn + +theorem ContDiffWithinAt.contDiffOn' (hm : m ≤ n) (h' : m = ∞ → n = ω) (h : ContDiffWithinAt 𝕜 n f s x) : ∃ u, IsOpen u ∧ x ∈ u ∧ ContDiffOn 𝕜 m f (insert x s ∩ u) := by - rcases h m hm with ⟨t, ht, p, hp⟩ - rcases mem_nhdsWithin.1 ht with ⟨u, huo, hxu, hut⟩ - rw [inter_comm] at hut - exact ⟨u, huo, hxu, (hp.mono hut).contDiffOn⟩ + rcases eq_or_ne n ω with rfl | hn + · obtain ⟨t, ht, p, hp, h'p⟩ := h + rcases mem_nhdsWithin.1 ht with ⟨u, huo, hxu, hut⟩ + rw [inter_comm] at hut + refine ⟨u, huo, hxu, ?_⟩ + suffices ContDiffOn 𝕜 ω f (insert x s ∩ u) from this.of_le le_top + intro y hy + refine ⟨insert x s ∩ u, ?_, p, hp.mono hut, fun i ↦ (h'p i).mono hut⟩ + simp only [insert_eq_of_mem, hy, self_mem_nhdsWithin] + · match m with + | ω => simp [hn] at hm + | ∞ => exact (hn (h' rfl)).elim + | (m : ℕ) => + rcases contDiffWithinAt_nat.1 (h.of_le hm) with ⟨t, ht, p, hp⟩ + rcases mem_nhdsWithin.1 ht with ⟨u, huo, hxu, hut⟩ + rw [inter_comm] at hut + exact ⟨u, huo, hxu, (hp.mono hut).contDiffOn⟩ -theorem ContDiffWithinAt.contDiffOn {m : ℕ} (hm : m ≤ n) (h : ContDiffWithinAt 𝕜 n f s x) : - ∃ u ∈ 𝓝[insert x s] x, u ⊆ insert x s ∧ ContDiffOn 𝕜 m f u := - let ⟨_u, uo, xu, h⟩ := h.contDiffOn' hm - ⟨_, inter_mem_nhdsWithin _ (uo.mem_nhds xu), inter_subset_left, h⟩ +theorem ContDiffWithinAt.contDiffOn (hm : m ≤ n) (h' : m = ∞ → n = ω) + (h : ContDiffWithinAt 𝕜 n f s x) : + ∃ u ∈ 𝓝[insert x s] x, u ⊆ insert x s ∧ ContDiffOn 𝕜 m f u := by + obtain ⟨_u, uo, xu, h⟩ := h.contDiffOn' hm h' + exact ⟨_, inter_mem_nhdsWithin _ (uo.mem_nhds xu), inter_subset_left, h⟩ + +theorem ContDiffOn.analyticOn (h : ContDiffOn 𝕜 ω f s) : AnalyticOn 𝕜 f s := + fun x hx ↦ (h x hx).analyticWithinAt /-- A function is `C^n` within a set at a point, for `n : ℕ`, if and only if it is `C^n` on a neighborhood of this point. -/ -theorem contDiffWithinAt_iff_contDiffOn_nhds {n : ℕ} : +theorem contDiffWithinAt_iff_contDiffOn_nhds (hn : n ≠ ∞) : ContDiffWithinAt 𝕜 n f s x ↔ ∃ u ∈ 𝓝[insert x s] x, ContDiffOn 𝕜 n f u := by refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ - · rcases h.contDiffOn le_rfl with ⟨u, hu, h'u⟩ + · rcases h.contDiffOn le_rfl (by simp [hn]) with ⟨u, hu, h'u⟩ exact ⟨u, hu, h'u.2⟩ · rcases h with ⟨u, u_mem, hu⟩ have : x ∈ u := mem_of_mem_nhdsWithin (mem_insert x s) u_mem exact (hu x this).mono_of_mem_nhdsWithin (nhdsWithin_mono _ (subset_insert x s) u_mem) -protected theorem ContDiffWithinAt.eventually {n : ℕ} (h : ContDiffWithinAt 𝕜 n f s x) : +protected theorem ContDiffWithinAt.eventually (h : ContDiffWithinAt 𝕜 n f s x) (hn : n ≠ ∞) : ∀ᶠ y in 𝓝[insert x s] x, ContDiffWithinAt 𝕜 n f s y := by - rcases h.contDiffOn le_rfl with ⟨u, hu, _, hd⟩ + rcases h.contDiffOn le_rfl (by simp [hn]) with ⟨u, hu, _, hd⟩ have : ∀ᶠ y : E in 𝓝[insert x s] x, u ∈ 𝓝[insert x s] y ∧ y ∈ u := (eventually_eventually_nhdsWithin.2 hu).and hu refine this.mono fun y hy => (hd y hy.2).mono_of_mem_nhdsWithin ?_ exact nhdsWithin_mono y (subset_insert _ _) hy.1 -theorem ContDiffOn.of_le (h : ContDiffOn 𝕜 n f s) (hmn : m ≤ n) : ContDiffOn 𝕜 m f s := fun x hx => - (h x hx).of_le hmn - -theorem ContDiffOn.of_succ {n : ℕ} (h : ContDiffOn 𝕜 (n + 1) f s) : ContDiffOn 𝕜 n f s := - h.of_le <| WithTop.coe_le_coe.mpr le_self_add +theorem ContDiffOn.of_succ (h : ContDiffOn 𝕜 (n + 1) f s) : ContDiffOn 𝕜 n f s := + h.of_le le_self_add -theorem ContDiffOn.one_of_succ {n : ℕ} (h : ContDiffOn 𝕜 (n + 1) f s) : ContDiffOn 𝕜 1 f s := - h.of_le <| WithTop.coe_le_coe.mpr le_add_self +theorem ContDiffOn.one_of_succ (h : ContDiffOn 𝕜 (n + 1) f s) : ContDiffOn 𝕜 1 f s := + h.of_le le_add_self theorem contDiffOn_iff_forall_nat_le {n : ℕ∞} : ContDiffOn 𝕜 n f s ↔ ∀ m : ℕ, ↑m ≤ n → ContDiffOn 𝕜 m f s := - ⟨fun H _ hm => H.of_le (mod_cast hm), fun H x hx m hm => H m (mod_cast hm) x hx m le_rfl⟩ + ⟨fun H _ hm => H.of_le (mod_cast hm), fun H x hx m hm => H m hm x hx m le_rfl⟩ -theorem contDiffOn_top : ContDiffOn 𝕜 ∞ f s ↔ ∀ n : ℕ, ContDiffOn 𝕜 n f s := +theorem contDiffOn_infty : ContDiffOn 𝕜 ∞ f s ↔ ∀ n : ℕ, ContDiffOn 𝕜 n f s := contDiffOn_iff_forall_nat_le.trans <| by simp only [le_top, forall_prop_of_true] +@[deprecated (since := "2024-11-27")] alias contDiffOn_top := contDiffOn_infty +@[deprecated (since := "2024-11-27")] +alias contDiffOn_infty_iff_contDiffOn_omega := contDiffOn_infty + theorem contDiffOn_all_iff_nat : (∀ (n : ℕ∞), ContDiffOn 𝕜 n f s) ↔ ∀ n : ℕ, ContDiffOn 𝕜 n f s := by refine ⟨fun H n => H n, ?_⟩ rintro H (_ | n) - exacts [contDiffOn_top.2 H, H n] + exacts [contDiffOn_infty.2 H, H n] theorem ContDiffOn.continuousOn (h : ContDiffOn 𝕜 n f s) : ContinuousOn f s := fun x hx => (h x hx).continuousWithinAt @@ -478,28 +603,28 @@ theorem contDiffOn_of_locally_contDiffOn exact IsOpen.mem_nhds u_open xu /-- A function is `C^(n + 1)` on a domain iff locally, it has a derivative which is `C^n`. -/ -theorem contDiffOn_succ_iff_hasFDerivWithinAt {n : ℕ} : +theorem contDiffOn_succ_iff_hasFDerivWithinAt (hn : n ≠ ∞) : ContDiffOn 𝕜 (n + 1) f s ↔ - ∀ x ∈ s, ∃ u ∈ 𝓝[insert x s] x, ∃ f' : E → E →L[𝕜] F, + ∀ x ∈ s, ∃ u ∈ 𝓝[insert x s] x, (n = ω → AnalyticOn 𝕜 f u) ∧ ∃ f' : E → E →L[𝕜] F, (∀ x ∈ u, HasFDerivWithinAt f (f' x) u x) ∧ ContDiffOn 𝕜 n f' u := by constructor · intro h x hx - rcases (h x hx) n.succ le_rfl with ⟨u, hu, p, Hp⟩ - refine - ⟨u, hu, fun y => (continuousMultilinearCurryFin1 𝕜 E F) (p y 1), fun y hy => - Hp.hasFDerivWithinAt (mod_cast (Nat.le_add_left 1 n)) hy, ?_⟩ - rw [show (n.succ : WithTop ℕ∞) = (n : ℕ) + 1 from rfl, - hasFTaylorSeriesUpToOn_succ_iff_right] at Hp - intro z hz m hm - refine ⟨u, ?_, fun x : E => (p x).shift, Hp.2.2.of_le (mod_cast hm)⟩ - -- Porting note: without the explicit arguments `convert` can not determine the type. - convert @self_mem_nhdsWithin _ _ z u - exact insert_eq_of_mem hz + rcases (contDiffWithinAt_succ_iff_hasFDerivWithinAt hn).1 (h x hx) with + ⟨u, hu, f_an, f', hf', Hf'⟩ + rcases Hf'.contDiffOn le_rfl (by simp [hn]) with ⟨v, vu, v'u, hv⟩ + rw [insert_eq_of_mem hx] at hu ⊢ + have xu : x ∈ u := mem_of_mem_nhdsWithin hx hu + rw [insert_eq_of_mem xu] at vu v'u + exact ⟨v, nhdsWithin_le_of_mem hu vu, fun h ↦ (f_an h).mono v'u, f', + fun y hy ↦ (hf' y (v'u hy)).mono v'u, hv⟩ · intro h x hx - rw [contDiffWithinAt_succ_iff_hasFDerivWithinAt] - rcases h x hx with ⟨u, u_nhbd, f', hu, hf'⟩ + rw [contDiffWithinAt_succ_iff_hasFDerivWithinAt hn] + rcases h x hx with ⟨u, u_nhbd, f_an, f', hu, hf'⟩ have : x ∈ u := mem_of_mem_nhdsWithin (mem_insert _ _) u_nhbd - exact ⟨u, u_nhbd, f', hu, hf' x this⟩ + exact ⟨u, u_nhbd, f_an, f', hu, hf' x this⟩ + + +/-! ### Iterated derivative within a set -/ @[simp] theorem contDiffOn_zero : ContDiffOn 𝕜 0 f s ↔ ContinuousOn f s := by @@ -526,7 +651,8 @@ theorem contDiffWithinAt_zero (hx : x ∈ s) : /-- When a function is `C^n` in a set `s` of unique differentiability, it admits `ftaylorSeriesWithin 𝕜 f s` as a Taylor series up to order `n` in `s`. -/ -protected theorem ContDiffOn.ftaylorSeriesWithin (h : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s) : +protected theorem ContDiffOn.ftaylorSeriesWithin + (h : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s) : HasFTaylorSeriesUpToOn n f (ftaylorSeriesWithin 𝕜 f s) s := by constructor · intro x _ @@ -556,7 +682,7 @@ protected theorem ContDiffOn.ftaylorSeriesWithin (h : ContDiffOn 𝕜 n f s) (hs · intro m hm apply continuousOn_of_locally_continuousOn intro x hx - rcases h x hx m (mod_cast hm) with ⟨u, hu, p, Hp⟩ + rcases (h x hx).of_le hm _ le_rfl with ⟨u, hu, p, Hp⟩ rcases mem_nhdsWithin.1 hu with ⟨o, o_open, xo, ho⟩ rw [insert_eq_of_mem hx] at ho rw [inter_comm] at ho @@ -568,6 +694,37 @@ protected theorem ContDiffOn.ftaylorSeriesWithin (h : ContDiffOn 𝕜 n f s) (hs exact (Hp.mono ho).eq_iteratedFDerivWithin_of_uniqueDiffOn le_rfl (hs.inter o_open) ⟨hy, yo⟩ exact ((Hp.mono ho).cont m le_rfl).congr fun y hy => (A y hy).symm +/-- On a set with unique differentiability, an analytic function is automatically `C^ω`, as its +successive derivatives are also analytic. This does not require completeness of the space. See +also `AnalyticOn.contDiffOn_of_completeSpace`.-/ +theorem AnalyticOn.contDiffOn (h : AnalyticOn 𝕜 f s) (hs : UniqueDiffOn 𝕜 s) : + ContDiffOn 𝕜 n f s := by + suffices ContDiffOn 𝕜 ω f s from this.of_le le_top + rcases h.exists_hasFTaylorSeriesUpToOn hs with ⟨p, hp⟩ + intro x hx + refine ⟨s, ?_, p, hp⟩ + rw [insert_eq_of_mem hx] + exact self_mem_nhdsWithin + +@[deprecated (since := "2024-09-26")] +alias AnalyticWithinOn.contDiffOn := AnalyticOn.contDiffOn + +/-- On a set with unique differentiability, an analytic function is automatically `C^ω`, as its +successive derivatives are also analytic. This does not require completeness of the space. See +also `AnalyticOnNhd.contDiffOn_of_completeSpace`. -/ +theorem AnalyticOnNhd.contDiffOn (h : AnalyticOnNhd 𝕜 f s) (hs : UniqueDiffOn 𝕜 s) : + ContDiffOn 𝕜 n f s := h.analyticOn.contDiffOn hs + +/-- An analytic function is automatically `C^ω` in a complete space -/ +theorem AnalyticOn.contDiffOn_of_completeSpace [CompleteSpace F] (h : AnalyticOn 𝕜 f s) : + ContDiffOn 𝕜 n f s := + fun x hx ↦ (h x hx).contDiffWithinAt + +/-- An analytic function is automatically `C^ω` in a complete space -/ +theorem AnalyticOnNhd.contDiffOn_of_completeSpace [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f s) : + ContDiffOn 𝕜 n f s := + h.analyticOn.contDiffOn_of_completeSpace + theorem contDiffOn_of_continuousOn_differentiableOn {n : ℕ∞} (Hcont : ∀ m : ℕ, m ≤ n → ContinuousOn (fun x => iteratedFDerivWithin 𝕜 m f s x) s) (Hdiff : ∀ m : ℕ, m < n → @@ -591,19 +748,47 @@ theorem contDiffOn_of_differentiableOn {n : ℕ∞} contDiffOn_of_continuousOn_differentiableOn (fun m hm => (h m hm).continuousOn) fun m hm => h m (le_of_lt hm) +theorem contDiffOn_of_analyticOn_iteratedFDerivWithin + (h : ∀ m, AnalyticOn 𝕜 (iteratedFDerivWithin 𝕜 m f s) s) : + ContDiffOn 𝕜 n f s := by + suffices ContDiffOn 𝕜 ω f s from this.of_le le_top + intro x hx + refine ⟨insert x s, self_mem_nhdsWithin, ftaylorSeriesWithin 𝕜 f s, ?_, ?_⟩ + · rw [insert_eq_of_mem hx] + constructor + · intro y _ + simp only [ftaylorSeriesWithin, ContinuousMultilinearMap.curry0_apply, + iteratedFDerivWithin_zero_apply] + · intro k _ y hy + exact ((h k).differentiableOn y hy).hasFDerivWithinAt + · intro k _ + exact (h k).continuousOn + · intro i + rw [insert_eq_of_mem hx] + exact h i + +theorem contDiffOn_omega_iff_analyticOn (hs : UniqueDiffOn 𝕜 s) : + ContDiffOn 𝕜 ω f s ↔ AnalyticOn 𝕜 f s := + ⟨fun h m ↦ h.analyticOn m, fun h ↦ h.contDiffOn hs⟩ + theorem ContDiffOn.continuousOn_iteratedFDerivWithin {m : ℕ} (h : ContDiffOn 𝕜 n f s) (hmn : m ≤ n) (hs : UniqueDiffOn 𝕜 s) : ContinuousOn (iteratedFDerivWithin 𝕜 m f s) s := - (h.ftaylorSeriesWithin hs).cont m (mod_cast hmn) + ((h.of_le hmn).ftaylorSeriesWithin hs).cont m le_rfl theorem ContDiffOn.differentiableOn_iteratedFDerivWithin {m : ℕ} (h : ContDiffOn 𝕜 n f s) (hmn : m < n) (hs : UniqueDiffOn 𝕜 s) : - DifferentiableOn 𝕜 (iteratedFDerivWithin 𝕜 m f s) s := fun x hx => - ((h.ftaylorSeriesWithin hs).fderivWithin m (mod_cast hmn) x hx).differentiableWithinAt + DifferentiableOn 𝕜 (iteratedFDerivWithin 𝕜 m f s) s := by + intro x hx + have : (m + 1 : ℕ) ≤ n := ENat.add_one_natCast_le_withTop_of_lt hmn + apply (((h.of_le this).ftaylorSeriesWithin hs).fderivWithin m ?_ x hx).differentiableWithinAt + exact_mod_cast lt_add_one m theorem ContDiffWithinAt.differentiableWithinAt_iteratedFDerivWithin {m : ℕ} (h : ContDiffWithinAt 𝕜 n f s x) (hmn : m < n) (hs : UniqueDiffOn 𝕜 (insert x s)) : DifferentiableWithinAt 𝕜 (iteratedFDerivWithin 𝕜 m f s) s x := by - rcases h.contDiffOn' (ENat.add_one_natCast_le_withTop_of_lt hmn) with ⟨u, uo, xu, hu⟩ + have : (m + 1 : WithTop ℕ∞) ≠ ∞ := Ne.symm (ne_of_beq_false rfl) + rcases h.contDiffOn' (ENat.add_one_natCast_le_withTop_of_lt hmn) (by simp [this]) + with ⟨u, uo, xu, hu⟩ set t := insert x s ∩ u have A : t =ᶠ[𝓝[≠] x] s := by simp only [set_eventuallyEq_iff_inf_principal, ← nhdsWithin_inter'] @@ -622,127 +807,140 @@ theorem contDiffOn_iff_continuousOn_differentiableOn {n : ℕ∞} (hs : UniqueDi ContDiffOn 𝕜 n f s ↔ (∀ m : ℕ, m ≤ n → ContinuousOn (fun x => iteratedFDerivWithin 𝕜 m f s x) s) ∧ ∀ m : ℕ, m < n → DifferentiableOn 𝕜 (fun x => iteratedFDerivWithin 𝕜 m f s x) s := - ⟨fun h => ⟨fun _m hm => h.continuousOn_iteratedFDerivWithin (mod_cast hm) hs, fun _m hm => - h.differentiableOn_iteratedFDerivWithin (mod_cast hm) hs⟩, + ⟨fun h => ⟨fun _m hm => h.continuousOn_iteratedFDerivWithin (mod_cast hm) hs, + fun _m hm => h.differentiableOn_iteratedFDerivWithin (mod_cast hm) hs⟩, fun h => contDiffOn_of_continuousOn_differentiableOn h.1 h.2⟩ -theorem contDiffOn_succ_of_fderivWithin {n : ℕ} (hf : DifferentiableOn 𝕜 f s) +theorem contDiffOn_nat_iff_continuousOn_differentiableOn {n : ℕ} (hs : UniqueDiffOn 𝕜 s) : + ContDiffOn 𝕜 n f s ↔ + (∀ m : ℕ, m ≤ n → ContinuousOn (fun x => iteratedFDerivWithin 𝕜 m f s x) s) ∧ + ∀ m : ℕ, m < n → DifferentiableOn 𝕜 (fun x => iteratedFDerivWithin 𝕜 m f s x) s := by + rw [show n = ((n : ℕ∞) : WithTop ℕ∞) from rfl, contDiffOn_iff_continuousOn_differentiableOn hs] + simp + +theorem contDiffOn_succ_of_fderivWithin (hf : DifferentiableOn 𝕜 f s) + (h' : n = ω → AnalyticOn 𝕜 f s) (h : ContDiffOn 𝕜 n (fun y => fderivWithin 𝕜 f s y) s) : ContDiffOn 𝕜 (n + 1) f s := by - intro x hx - rw [contDiffWithinAt_succ_iff_hasFDerivWithinAt, insert_eq_of_mem hx] - exact - ⟨s, self_mem_nhdsWithin, fderivWithin 𝕜 f s, fun y hy => (hf y hy).hasFDerivWithinAt, h x hx⟩ + rcases eq_or_ne n ∞ with rfl | hn + · rw [ENat.coe_top_add_one, contDiffOn_infty] + intro m x hx + apply ContDiffWithinAt.of_le _ (show (m : WithTop ℕ∞) ≤ m + 1 from le_self_add) + rw [contDiffWithinAt_succ_iff_hasFDerivWithinAt (by simp), + insert_eq_of_mem hx] + exact ⟨s, self_mem_nhdsWithin, (by simp), fderivWithin 𝕜 f s, + fun y hy => (hf y hy).hasFDerivWithinAt, (h x hx).of_le (mod_cast le_top)⟩ + · intro x hx + rw [contDiffWithinAt_succ_iff_hasFDerivWithinAt hn, + insert_eq_of_mem hx] + exact ⟨s, self_mem_nhdsWithin, h', fderivWithin 𝕜 f s, + fun y hy => (hf y hy).hasFDerivWithinAt, h x hx⟩ + +theorem contDiffOn_of_analyticOn_of_fderivWithin (hf : AnalyticOn 𝕜 f s) + (h : ContDiffOn 𝕜 ω (fun y ↦ fderivWithin 𝕜 f s y) s) : ContDiffOn 𝕜 n f s := by + suffices ContDiffOn 𝕜 (ω + 1) f s from this.of_le le_top + exact contDiffOn_succ_of_fderivWithin hf.differentiableOn (fun _ ↦ hf) h /-- A function is `C^(n + 1)` on a domain with unique derivatives if and only if it is differentiable there, and its derivative (expressed with `fderivWithin`) is `C^n`. -/ -theorem contDiffOn_succ_iff_fderivWithin {n : ℕ} (hs : UniqueDiffOn 𝕜 s) : - ContDiffOn 𝕜 (n + 1) f s ↔ - DifferentiableOn 𝕜 f s ∧ ContDiffOn 𝕜 n (fun y => fderivWithin 𝕜 f s y) s := by - refine ⟨fun H => ?_, fun h => contDiffOn_succ_of_fderivWithin h.1 h.2⟩ - refine ⟨H.differentiableOn le_add_self, fun x hx => ?_⟩ - rcases contDiffWithinAt_succ_iff_hasFDerivWithinAt.1 (H x hx) with ⟨u, hu, f', hff', hf'⟩ - rcases mem_nhdsWithin.1 hu with ⟨o, o_open, xo, ho⟩ - rw [inter_comm, insert_eq_of_mem hx] at ho - have := hf'.mono ho - rw [contDiffWithinAt_inter' (mem_nhdsWithin_of_mem_nhds (IsOpen.mem_nhds o_open xo))] at this - apply this.congr_of_eventuallyEq_of_mem _ hx - have : o ∩ s ∈ 𝓝[s] x := mem_nhdsWithin.2 ⟨o, o_open, xo, Subset.refl _⟩ - rw [inter_comm] at this - refine Filter.eventuallyEq_of_mem this fun y hy => ?_ - have A : fderivWithin 𝕜 f (s ∩ o) y = f' y := - ((hff' y (ho hy)).mono ho).fderivWithin (hs.inter o_open y hy) - rwa [fderivWithin_inter (o_open.mem_nhds hy.2)] at A - -theorem contDiffOn_succ_iff_hasFDerivWithin {n : ℕ} (hs : UniqueDiffOn 𝕜 s) : +theorem contDiffOn_succ_iff_fderivWithin (hs : UniqueDiffOn 𝕜 s) : ContDiffOn 𝕜 (n + 1) f s ↔ + DifferentiableOn 𝕜 f s ∧ (n = ω → AnalyticOn 𝕜 f s) ∧ + ContDiffOn 𝕜 n (fderivWithin 𝕜 f s) s := by + refine ⟨fun H => ?_, fun h => contDiffOn_succ_of_fderivWithin h.1 h.2.1 h.2.2⟩ + refine ⟨H.differentiableOn le_add_self, ?_, fun x hx => ?_⟩ + · rintro rfl + exact H.analyticOn + have A (m : ℕ) (hm : m ≤ n) : ContDiffWithinAt 𝕜 m (fun y => fderivWithin 𝕜 f s y) s x := by + rcases (contDiffWithinAt_succ_iff_hasFDerivWithinAt (n := m) (ne_of_beq_false rfl)).1 + (H.of_le (add_le_add_right hm 1) x hx) with ⟨u, hu, -, f', hff', hf'⟩ + rcases mem_nhdsWithin.1 hu with ⟨o, o_open, xo, ho⟩ + rw [inter_comm, insert_eq_of_mem hx] at ho + have := hf'.mono ho + rw [contDiffWithinAt_inter' (mem_nhdsWithin_of_mem_nhds (IsOpen.mem_nhds o_open xo))] at this + apply this.congr_of_eventuallyEq_of_mem _ hx + have : o ∩ s ∈ 𝓝[s] x := mem_nhdsWithin.2 ⟨o, o_open, xo, Subset.refl _⟩ + rw [inter_comm] at this + refine Filter.eventuallyEq_of_mem this fun y hy => ?_ + have A : fderivWithin 𝕜 f (s ∩ o) y = f' y := + ((hff' y (ho hy)).mono ho).fderivWithin (hs.inter o_open y hy) + rwa [fderivWithin_inter (o_open.mem_nhds hy.2)] at A + match n with + | ω => exact (H.analyticOn.fderivWithin hs).contDiffOn hs (n := ω) x hx + | ∞ => exact contDiffWithinAt_infty.2 (fun m ↦ A m (mod_cast le_top)) + | (n : ℕ) => exact A n le_rfl + +theorem contDiffOn_succ_iff_hasFDerivWithinAt_of_uniqueDiffOn (hs : UniqueDiffOn 𝕜 s) : + ContDiffOn 𝕜 (n + 1) f s ↔ (n = ω → AnalyticOn 𝕜 f s) ∧ ∃ f' : E → E →L[𝕜] F, ContDiffOn 𝕜 n f' s ∧ ∀ x, x ∈ s → HasFDerivWithinAt f (f' x) s x := by rw [contDiffOn_succ_iff_fderivWithin hs] - refine ⟨fun h => ⟨fderivWithin 𝕜 f s, h.2, fun x hx => (h.1 x hx).hasFDerivWithinAt⟩, fun h => ?_⟩ + refine ⟨fun h => ⟨h.2.1, fderivWithin 𝕜 f s, h.2.2, + fun x hx => (h.1 x hx).hasFDerivWithinAt⟩, fun ⟨f_an, h⟩ => ?_⟩ rcases h with ⟨f', h1, h2⟩ - refine ⟨fun x hx => (h2 x hx).differentiableWithinAt, fun x hx => ?_⟩ + refine ⟨fun x hx => (h2 x hx).differentiableWithinAt, f_an, fun x hx => ?_⟩ exact (h1 x hx).congr_of_mem (fun y hy => (h2 y hy).fderivWithin (hs y hy)) hx +@[deprecated (since := "2024-11-27")] +alias contDiffOn_succ_iff_hasFDerivWithin := contDiffOn_succ_iff_hasFDerivWithinAt_of_uniqueDiffOn + +theorem contDiffOn_infty_iff_fderivWithin (hs : UniqueDiffOn 𝕜 s) : + ContDiffOn 𝕜 ∞ f s ↔ DifferentiableOn 𝕜 f s ∧ ContDiffOn 𝕜 ∞ (fderivWithin 𝕜 f s) s := by + rw [show ∞ = ∞ + 1 from rfl, contDiffOn_succ_iff_fderivWithin hs] + simp + +@[deprecated (since := "2024-11-27")] +alias contDiffOn_top_iff_fderivWithin := contDiffOn_infty_iff_fderivWithin + /-- A function is `C^(n + 1)` on an open domain if and only if it is differentiable there, and its derivative (expressed with `fderiv`) is `C^n`. -/ -theorem contDiffOn_succ_iff_fderiv_of_isOpen {n : ℕ} (hs : IsOpen s) : +theorem contDiffOn_succ_iff_fderiv_of_isOpen (hs : IsOpen s) : ContDiffOn 𝕜 (n + 1) f s ↔ - DifferentiableOn 𝕜 f s ∧ ContDiffOn 𝕜 n (fun y => fderiv 𝕜 f y) s := by + DifferentiableOn 𝕜 f s ∧ (n = ω → AnalyticOn 𝕜 f s) ∧ + ContDiffOn 𝕜 n (fderiv 𝕜 f) s := by rw [contDiffOn_succ_iff_fderivWithin hs.uniqueDiffOn] - exact Iff.rfl.and (contDiffOn_congr fun x hx ↦ fderivWithin_of_isOpen hs hx) + exact Iff.rfl.and (Iff.rfl.and (contDiffOn_congr fun x hx ↦ fderivWithin_of_isOpen hs hx)) -/-- A function is `C^∞` on a domain with unique derivatives if and only if it is differentiable -there, and its derivative (expressed with `fderivWithin`) is `C^∞`. -/ -theorem contDiffOn_top_iff_fderivWithin (hs : UniqueDiffOn 𝕜 s) : - ContDiffOn 𝕜 ∞ f s ↔ - DifferentiableOn 𝕜 f s ∧ ContDiffOn 𝕜 ∞ (fun y => fderivWithin 𝕜 f s y) s := by - constructor - · intro h - refine ⟨h.differentiableOn (mod_cast le_top), ?_⟩ - refine contDiffOn_top.2 fun n => ((contDiffOn_succ_iff_fderivWithin hs).1 ?_).2 - exact h.of_le (mod_cast le_top) - · intro h - refine contDiffOn_top.2 fun n => ?_ - have A : (n : ℕ∞) ≤ ∞ := mod_cast le_top - apply ((contDiffOn_succ_iff_fderivWithin hs).2 ⟨h.1, h.2.of_le A⟩).of_le - exact_mod_cast (Nat.le_succ n) - -/-- A function is `C^∞` on an open domain if and only if it is differentiable there, and its -derivative (expressed with `fderiv`) is `C^∞`. -/ -theorem contDiffOn_top_iff_fderiv_of_isOpen (hs : IsOpen s) : - ContDiffOn 𝕜 ∞ f s ↔ DifferentiableOn 𝕜 f s ∧ ContDiffOn 𝕜 ∞ (fun y => fderiv 𝕜 f y) s := by - rw [contDiffOn_top_iff_fderivWithin hs.uniqueDiffOn] - exact Iff.rfl.and <| contDiffOn_congr fun x hx ↦ fderivWithin_of_isOpen hs hx +theorem contDiffOn_infty_iff_fderiv_of_isOpen (hs : IsOpen s) : + ContDiffOn 𝕜 ∞ f s ↔ DifferentiableOn 𝕜 f s ∧ ContDiffOn 𝕜 ∞ (fderiv 𝕜 f) s := by + rw [show ∞ = ∞ + 1 from rfl, contDiffOn_succ_iff_fderiv_of_isOpen hs] + simp + +@[deprecated (since := "2024-11-27")] +alias contDiffOn_top_iff_fderiv_of_isOpen := contDiffOn_infty_iff_fderiv_of_isOpen protected theorem ContDiffOn.fderivWithin (hf : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s) - (hmn : m + 1 ≤ n) : ContDiffOn 𝕜 m (fun y => fderivWithin 𝕜 f s y) s := by - rcases le_or_lt ∞ n with hn | hn - · have : ContDiffOn 𝕜 ∞ (fun y ↦ fderivWithin 𝕜 f s y) s := - ((contDiffOn_top_iff_fderivWithin hs).1 (hf.of_le hn)).2 - intro x hx k hk - exact this x hx k (mod_cast le_top) - · match m with - | ω => simpa using hmn.trans_lt hn - | ∞ => simpa using hmn.trans_lt hn - | (m : ℕ) => - change (m.succ : ℕ∞) ≤ n at hmn - exact ((contDiffOn_succ_iff_fderivWithin hs).1 (hf.of_le hmn)).2 + (hmn : m + 1 ≤ n) : ContDiffOn 𝕜 m (fderivWithin 𝕜 f s) s := + ((contDiffOn_succ_iff_fderivWithin hs).1 (hf.of_le hmn)).2.2 theorem ContDiffOn.fderiv_of_isOpen (hf : ContDiffOn 𝕜 n f s) (hs : IsOpen s) (hmn : m + 1 ≤ n) : - ContDiffOn 𝕜 m (fun y => fderiv 𝕜 f y) s := + ContDiffOn 𝕜 m (fderiv 𝕜 f) s := (hf.fderivWithin hs.uniqueDiffOn hmn).congr fun _ hx => (fderivWithin_of_isOpen hs hx).symm theorem ContDiffOn.continuousOn_fderivWithin (h : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s) - (hn : 1 ≤ n) : ContinuousOn (fun x => fderivWithin 𝕜 f s x) s := - ((contDiffOn_succ_iff_fderivWithin hs).1 (h.of_le hn)).2.continuousOn + (hn : 1 ≤ n) : ContinuousOn (fderivWithin 𝕜 f s) s := + ((contDiffOn_succ_iff_fderivWithin hs).1 + (h.of_le (show 0 + (1 : WithTop ℕ∞) ≤ n from hn))).2.2.continuousOn theorem ContDiffOn.continuousOn_fderiv_of_isOpen (h : ContDiffOn 𝕜 n f s) (hs : IsOpen s) - (hn : 1 ≤ n) : ContinuousOn (fun x => fderiv 𝕜 f x) s := - ((contDiffOn_succ_iff_fderiv_of_isOpen hs).1 (h.of_le hn)).2.continuousOn - -/-- The following lemma will be removed when the definition of `C^ω` will be corrected. For now, -it is only there as a convenient shortcut. -/ -theorem contDiffOn_infty_iff_contDiffOn_omega : - ContDiffOn 𝕜 ∞ f s ↔ ContDiffOn 𝕜 ω f s := by - have A (m : ℕ) : m ≤ ∞ := mod_cast le_top - simp [ContDiffOn, ContDiffWithinAt, hasFTaylorSeriesUpTo_top_iff, A] + (hn : 1 ≤ n) : ContinuousOn (fderiv 𝕜 f) s := + ((contDiffOn_succ_iff_fderiv_of_isOpen hs).1 + (h.of_le (show 0 + (1 : WithTop ℕ∞) ≤ n from hn))).2.2.continuousOn /-! ### Smooth functions at a point -/ -variable (𝕜) - +variable (𝕜) in /-- A function is continuously differentiable up to `n` at a point `x` if, for any integer `k ≤ n`, there is a neighborhood of `x` where `f` admits derivatives up to order `n`, which are continuous. -/ def ContDiffAt (n : WithTop ℕ∞) (f : E → F) (x : E) : Prop := ContDiffWithinAt 𝕜 n f univ x -variable {𝕜} - theorem contDiffWithinAt_univ : ContDiffWithinAt 𝕜 n f univ x ↔ ContDiffAt 𝕜 n f x := Iff.rfl -theorem contDiffAt_top : ContDiffAt 𝕜 ∞ f x ↔ ∀ n : ℕ, ContDiffAt 𝕜 n f x := by - simp [← contDiffWithinAt_univ, contDiffWithinAt_top] +theorem contDiffAt_infty : ContDiffAt 𝕜 ∞ f x ↔ ∀ n : ℕ, ContDiffAt 𝕜 n f x := by + simp [← contDiffWithinAt_univ, contDiffWithinAt_infty] + +@[deprecated (since := "2024-11-27")] alias contDiffAt_top := contDiffAt_infty theorem ContDiffAt.contDiffWithinAt (h : ContDiffAt 𝕜 n f x) : ContDiffWithinAt 𝕜 n f s x := h.mono (subset_univ _) @@ -772,6 +970,20 @@ theorem ContDiffAt.of_le (h : ContDiffAt 𝕜 n f x) (hmn : m ≤ n) : ContDiffA theorem ContDiffAt.continuousAt (h : ContDiffAt 𝕜 n f x) : ContinuousAt f x := by simpa [continuousWithinAt_univ] using h.continuousWithinAt +theorem ContDiffAt.analyticAt (h : ContDiffAt 𝕜 ω f x) : AnalyticAt 𝕜 f x := by + rw [← contDiffWithinAt_univ] at h + rw [← analyticWithinAt_univ] + exact h.analyticWithinAt + +/-- In a complete space, a function which is analytic at a point is also `C^ω` there. +Note that the same statement for `AnalyticOn` does not require completeness, see +`AnalyticOn.contDiffOn`. -/ +theorem AnalyticAt.contDiffAt [CompleteSpace F] (h : AnalyticAt 𝕜 f x) : + ContDiffAt 𝕜 n f x := by + rw [← contDiffWithinAt_univ] + rw [← analyticWithinAt_univ] at h + exact h.contDiffWithinAt + @[simp] theorem contDiffWithinAt_compl_self : ContDiffWithinAt 𝕜 n f {x}ᶜ x ↔ ContDiffAt 𝕜 n f x := by @@ -782,18 +994,18 @@ theorem ContDiffAt.differentiableAt (h : ContDiffAt 𝕜 n f x) (hn : 1 ≤ n) : DifferentiableAt 𝕜 f x := by simpa [hn, differentiableWithinAt_univ] using h.differentiableWithinAt -nonrec lemma ContDiffAt.contDiffOn {m : ℕ} (h : ContDiffAt 𝕜 n f x) (hm : m ≤ n) : +nonrec lemma ContDiffAt.contDiffOn (h : ContDiffAt 𝕜 n f x) (hm : m ≤ n) (h' : m = ∞ → n = ω): ∃ u ∈ 𝓝 x, ContDiffOn 𝕜 m f u := by - simpa [nhdsWithin_univ] using h.contDiffOn hm + simpa [nhdsWithin_univ] using h.contDiffOn hm h' /-- A function is `C^(n + 1)` at a point iff locally, it has a derivative which is `C^n`. -/ theorem contDiffAt_succ_iff_hasFDerivAt {n : ℕ} : - ContDiffAt 𝕜 (n + 1) f x ↔ - ∃ f' : E → E →L[𝕜] F, (∃ u ∈ 𝓝 x, ∀ x ∈ u, HasFDerivAt f (f' x) x) ∧ ContDiffAt 𝕜 n f' x := by - rw [← contDiffWithinAt_univ, contDiffWithinAt_succ_iff_hasFDerivWithinAt] + ContDiffAt 𝕜 (n + 1) f x ↔ ∃ f' : E → E →L[𝕜] F, + (∃ u ∈ 𝓝 x, ∀ x ∈ u, HasFDerivAt f (f' x) x) ∧ ContDiffAt 𝕜 n f' x := by + rw [← contDiffWithinAt_univ, contDiffWithinAt_succ_iff_hasFDerivWithinAt (by simp)] simp only [nhdsWithin_univ, exists_prop, mem_univ, insert_eq_of_mem] constructor - · rintro ⟨u, H, f', h_fderiv, h_cont_diff⟩ + · rintro ⟨u, H, -, f', h_fderiv, h_cont_diff⟩ rcases mem_nhds_iff.mp H with ⟨t, htu, ht, hxt⟩ refine ⟨f', ⟨t, ?_⟩, h_cont_diff.contDiffAt H⟩ refine ⟨mem_nhds_iff.mpr ⟨t, Subset.rfl, ht, hxt⟩, ?_⟩ @@ -801,41 +1013,53 @@ theorem contDiffAt_succ_iff_hasFDerivAt {n : ℕ} : refine (h_fderiv y (htu hyt)).hasFDerivAt ?_ exact mem_nhds_iff.mpr ⟨t, htu, ht, hyt⟩ · rintro ⟨f', ⟨u, H, h_fderiv⟩, h_cont_diff⟩ - refine ⟨u, H, f', ?_, h_cont_diff.contDiffWithinAt⟩ - intro x hxu + refine ⟨u, H, by simp, f', fun x hxu ↦ ?_, h_cont_diff.contDiffWithinAt⟩ exact (h_fderiv x hxu).hasFDerivWithinAt -protected theorem ContDiffAt.eventually {n : ℕ} (h : ContDiffAt 𝕜 n f x) : +protected theorem ContDiffAt.eventually (h : ContDiffAt 𝕜 n f x) (h' : n ≠ ∞) : ∀ᶠ y in 𝓝 x, ContDiffAt 𝕜 n f y := by - simpa [nhdsWithin_univ] using ContDiffWithinAt.eventually h + simpa [nhdsWithin_univ] using ContDiffWithinAt.eventually h h' /-! ### Smooth functions -/ -variable (𝕜) - +variable (𝕜) in /-- A function is continuously differentiable up to `n` if it admits derivatives up to order `n`, which are continuous. Contrary to the case of definitions in domains (where derivatives might not be unique) we do not need to localize the definition in space or time. -/ def ContDiff (n : WithTop ℕ∞) (f : E → F) : Prop := - ∃ p : E → FormalMultilinearSeries 𝕜 E F, HasFTaylorSeriesUpTo n f p - -variable {𝕜} + match n with + | ω => ∃ p : E → FormalMultilinearSeries 𝕜 E F, HasFTaylorSeriesUpTo ⊤ f p + ∧ ∀ i, AnalyticOnNhd 𝕜 (fun x ↦ p x i) univ + | (n : ℕ∞) => ∃ p : E → FormalMultilinearSeries 𝕜 E F, HasFTaylorSeriesUpTo n f p /-- If `f` has a Taylor series up to `n`, then it is `C^n`. -/ -theorem HasFTaylorSeriesUpTo.contDiff {f' : E → FormalMultilinearSeries 𝕜 E F} +theorem HasFTaylorSeriesUpTo.contDiff {n : ℕ∞} {f' : E → FormalMultilinearSeries 𝕜 E F} (hf : HasFTaylorSeriesUpTo n f f') : ContDiff 𝕜 n f := ⟨f', hf⟩ theorem contDiffOn_univ : ContDiffOn 𝕜 n f univ ↔ ContDiff 𝕜 n f := by - constructor - · intro H - use ftaylorSeriesWithin 𝕜 f univ - rw [← hasFTaylorSeriesUpToOn_univ_iff] - exact H.ftaylorSeriesWithin uniqueDiffOn_univ - · rintro ⟨p, hp⟩ x _ m hm - exact ⟨univ, Filter.univ_sets _, p, (hp.hasFTaylorSeriesUpToOn univ).of_le - (mod_cast hm)⟩ + match n with + | ω => + constructor + · intro H + use ftaylorSeriesWithin 𝕜 f univ + rw [← hasFTaylorSeriesUpToOn_univ_iff] + refine ⟨H.ftaylorSeriesWithin uniqueDiffOn_univ, fun i ↦ ?_⟩ + rw [← analyticOn_univ] + exact H.analyticOn.iteratedFDerivWithin uniqueDiffOn_univ _ + · rintro ⟨p, hp, h'p⟩ x _ + exact ⟨univ, Filter.univ_sets _, p, (hp.hasFTaylorSeriesUpToOn univ).of_le le_top, + fun i ↦ (h'p i).analyticOn⟩ + | (n : ℕ∞) => + constructor + · intro H + use ftaylorSeriesWithin 𝕜 f univ + rw [← hasFTaylorSeriesUpToOn_univ_iff] + exact H.ftaylorSeriesWithin uniqueDiffOn_univ + · rintro ⟨p, hp⟩ x _ m hm + exact ⟨univ, Filter.univ_sets _, p, + (hp.hasFTaylorSeriesUpToOn univ).of_le (mod_cast hm)⟩ theorem contDiff_iff_contDiffAt : ContDiff 𝕜 n f ↔ ∀ x, ContDiffAt 𝕜 n f x := by simp [← contDiffOn_univ, ContDiffOn, ContDiffAt] @@ -846,16 +1070,14 @@ theorem ContDiff.contDiffAt (h : ContDiff 𝕜 n f) : ContDiffAt 𝕜 n f x := theorem ContDiff.contDiffWithinAt (h : ContDiff 𝕜 n f) : ContDiffWithinAt 𝕜 n f s x := h.contDiffAt.contDiffWithinAt -/-- The following lemma will be removed when the definition of `C^ω` will be corrected. For now, -it is only there as a convenient shortcut. -/ -theorem contDiff_infty_iff_contDiff_omega : - ContDiff 𝕜 ∞ f ↔ ContDiff 𝕜 ω f := by - simp [ContDiff, hasFTaylorSeriesUpTo_top_iff] +theorem contDiff_infty : ContDiff 𝕜 ∞ f ↔ ∀ n : ℕ, ContDiff 𝕜 n f := by + simp [contDiffOn_univ.symm, contDiffOn_infty] + +@[deprecated (since := "2024-11-25")] alias contDiff_top := contDiff_infty -theorem contDiff_top : ContDiff 𝕜 ∞ f ↔ ∀ n : ℕ, ContDiff 𝕜 n f := by - simp [contDiffOn_univ.symm, contDiffOn_top] +@[deprecated (since := "2024-11-25")] alias contDiff_infty_iff_contDiff_omega := contDiff_infty -theorem contDiff_all_iff_nat : (∀ (n : ℕ∞), ContDiff 𝕜 n f) ↔ ∀ n : ℕ, ContDiff 𝕜 n f := by +theorem contDiff_all_iff_nat : (∀ n : ℕ∞, ContDiff 𝕜 n f) ↔ ∀ n : ℕ, ContDiff 𝕜 n f := by simp only [← contDiffOn_univ, contDiffOn_all_iff_nat] theorem ContDiff.contDiffOn (h : ContDiff 𝕜 n f) : ContDiffOn 𝕜 n f s := @@ -872,18 +1094,18 @@ theorem contDiffAt_zero : ContDiffAt 𝕜 0 f x ↔ ∃ u ∈ 𝓝 x, Continuous theorem contDiffAt_one_iff : ContDiffAt 𝕜 1 f x ↔ ∃ f' : E → E →L[𝕜] F, ∃ u ∈ 𝓝 x, ContinuousOn f' u ∧ ∀ x ∈ u, HasFDerivAt f (f' x) x := by - rw [show (1 : WithTop ℕ∞) = (0 : ℕ) + 1 from rfl, contDiffAt_succ_iff_hasFDerivAt] - simp_rw [show ((0 : ℕ) : WithTop ℕ∞) = 0 from rfl, contDiffAt_zero, - exists_mem_and_iff antitone_bforall antitone_continuousOn, and_comm] + rw [show (1 : WithTop ℕ∞) = (0 : ℕ) + 1 from rfl] + simp_rw [contDiffAt_succ_iff_hasFDerivAt, show ((0 : ℕ) : WithTop ℕ∞) = 0 from rfl, + contDiffAt_zero, exists_mem_and_iff antitone_bforall antitone_continuousOn, and_comm] theorem ContDiff.of_le (h : ContDiff 𝕜 n f) (hmn : m ≤ n) : ContDiff 𝕜 m f := contDiffOn_univ.1 <| (contDiffOn_univ.2 h).of_le hmn -theorem ContDiff.of_succ {n : ℕ} (h : ContDiff 𝕜 (n + 1) f) : ContDiff 𝕜 n f := - h.of_le <| WithTop.coe_le_coe.mpr le_self_add +theorem ContDiff.of_succ (h : ContDiff 𝕜 (n + 1) f) : ContDiff 𝕜 n f := + h.of_le le_self_add -theorem ContDiff.one_of_succ {n : ℕ} (h : ContDiff 𝕜 (n + 1) f) : ContDiff 𝕜 1 f := - h.of_le <| WithTop.coe_le_coe.mpr le_add_self +theorem ContDiff.one_of_succ (h : ContDiff 𝕜 (n + 1) f) : ContDiff 𝕜 1 f := by + apply h.of_le le_add_self theorem ContDiff.continuous (h : ContDiff 𝕜 n f) : Continuous f := contDiff_zero.1 (h.of_le bot_le) @@ -900,16 +1122,37 @@ theorem contDiff_iff_forall_nat_le {n : ℕ∞} : theorem contDiff_succ_iff_hasFDerivAt {n : ℕ} : ContDiff 𝕜 (n + 1) f ↔ ∃ f' : E → E →L[𝕜] F, ContDiff 𝕜 n f' ∧ ∀ x, HasFDerivAt f (f' x) x := by - simp only [← contDiffOn_univ, ← hasFDerivWithinAt_univ, - contDiffOn_succ_iff_hasFDerivWithin uniqueDiffOn_univ, Set.mem_univ, forall_true_left] + simp only [← contDiffOn_univ, ← hasFDerivWithinAt_univ, Set.mem_univ, forall_true_left, + contDiffOn_succ_iff_hasFDerivWithinAt_of_uniqueDiffOn uniqueDiffOn_univ, + WithTop.natCast_ne_top, analyticOn_univ, false_implies, true_and] theorem contDiff_one_iff_hasFDerivAt : ContDiff 𝕜 1 f ↔ ∃ f' : E → E →L[𝕜] F, Continuous f' ∧ ∀ x, HasFDerivAt f (f' x) x := by convert contDiff_succ_iff_hasFDerivAt using 4; simp +theorem AnalyticOn.contDiff (hf : AnalyticOn 𝕜 f univ) : ContDiff 𝕜 n f := by + rw [← contDiffOn_univ] + exact hf.contDiffOn (n := n) uniqueDiffOn_univ + +theorem AnalyticOnNhd.contDiff (hf : AnalyticOnNhd 𝕜 f univ) : ContDiff 𝕜 n f := + hf.analyticOn.contDiff + +theorem ContDiff.analyticOnNhd (h : ContDiff 𝕜 ω f) : AnalyticOnNhd 𝕜 f s := by + rw [← contDiffOn_univ] at h + have := h.analyticOn + rw [analyticOn_univ] at this + exact this.mono (subset_univ _) + +theorem contDiff_omega_iff_analyticOnNhd : + ContDiff 𝕜 ω f ↔ AnalyticOnNhd 𝕜 f univ := + ⟨fun h ↦ h.analyticOnNhd, fun h ↦ h.contDiff⟩ + +/-! ### Iterated derivative -/ + + /-- When a function is `C^n` in a set `s` of unique differentiability, it admits `ftaylorSeriesWithin 𝕜 f s` as a Taylor series up to order `n` in `s`. -/ -theorem contDiff_iff_ftaylorSeries : +theorem contDiff_iff_ftaylorSeries {n : ℕ∞} : ContDiff 𝕜 n f ↔ HasFTaylorSeriesUpTo n f (ftaylorSeries 𝕜 f) := by constructor · rw [← contDiffOn_univ, ← hasFTaylorSeriesUpToOn_univ_iff, ← ftaylorSeriesWithin_univ] @@ -923,6 +1166,13 @@ theorem contDiff_iff_continuous_differentiable {n : ℕ∞} : simp [contDiffOn_univ.symm, continuous_iff_continuousOn_univ, differentiableOn_univ.symm, iteratedFDerivWithin_univ, contDiffOn_iff_continuousOn_differentiableOn uniqueDiffOn_univ] +theorem contDiff_nat_iff_continuous_differentiable {n : ℕ} : + ContDiff 𝕜 n f ↔ + (∀ m : ℕ, m ≤ n → Continuous fun x => iteratedFDeriv 𝕜 m f x) ∧ + ∀ m : ℕ, m < n → Differentiable 𝕜 fun x => iteratedFDeriv 𝕜 m f x := by + rw [show n = ((n : ℕ∞) : WithTop ℕ∞) from rfl, contDiff_iff_continuous_differentiable] + simp + /-- If `f` is `C^n` then its `m`-times iterated derivative is continuous for `m ≤ n`. -/ theorem ContDiff.continuous_iteratedFDeriv {m : ℕ} (hm : m ≤ n) (hf : ContDiff 𝕜 n f) : Continuous fun x => iteratedFDeriv 𝕜 m f x := @@ -941,24 +1191,27 @@ theorem contDiff_of_differentiable_iteratedFDeriv {n : ℕ∞} /-- A function is `C^(n + 1)` if and only if it is differentiable, and its derivative (formulated in terms of `fderiv`) is `C^n`. -/ -theorem contDiff_succ_iff_fderiv {n : ℕ} : - ContDiff 𝕜 (n + 1) f ↔ Differentiable 𝕜 f ∧ ContDiff 𝕜 n fun y => fderiv 𝕜 f y := by +theorem contDiff_succ_iff_fderiv : + ContDiff 𝕜 (n + 1) f ↔ Differentiable 𝕜 f ∧ (n = ω → AnalyticOnNhd 𝕜 f univ) ∧ + ContDiff 𝕜 n (fderiv 𝕜 f) := by simp only [← contDiffOn_univ, ← differentiableOn_univ, ← fderivWithin_univ, - contDiffOn_succ_iff_fderivWithin uniqueDiffOn_univ] + contDiffOn_succ_iff_fderivWithin uniqueDiffOn_univ, analyticOn_univ] + +theorem contDiff_one_iff_fderiv : + ContDiff 𝕜 1 f ↔ Differentiable 𝕜 f ∧ Continuous (fderiv 𝕜 f) := by + rw [show (1 : WithTop ℕ∞) = 0 + 1 from rfl, contDiff_succ_iff_fderiv] + simp -theorem contDiff_one_iff_fderiv : ContDiff 𝕜 1 f ↔ Differentiable 𝕜 f ∧ Continuous (fderiv 𝕜 f) := - contDiff_succ_iff_fderiv.trans <| Iff.rfl.and contDiff_zero +theorem contDiff_infty_iff_fderiv : + ContDiff 𝕜 ∞ f ↔ Differentiable 𝕜 f ∧ ContDiff 𝕜 ∞ (fderiv 𝕜 f) := by + rw [show ∞ = ∞ + 1 from rfl, contDiff_succ_iff_fderiv] + simp -/-- A function is `C^∞` if and only if it is differentiable, -and its derivative (formulated in terms of `fderiv`) is `C^∞`. -/ -theorem contDiff_top_iff_fderiv : - ContDiff 𝕜 ∞ f ↔ Differentiable 𝕜 f ∧ ContDiff 𝕜 ∞ fun y => fderiv 𝕜 f y := by - simp only [← contDiffOn_univ, ← differentiableOn_univ, ← fderivWithin_univ] - rw [contDiffOn_top_iff_fderivWithin uniqueDiffOn_univ] +@[deprecated (since := "2024-11-27")] alias contDiff_top_iff_fderiv := contDiff_infty_iff_fderiv theorem ContDiff.continuous_fderiv (h : ContDiff 𝕜 n f) (hn : 1 ≤ n) : - Continuous fun x => fderiv 𝕜 f x := - (contDiff_succ_iff_fderiv.1 (h.of_le hn)).2.continuous + Continuous (fderiv 𝕜 f) := + (contDiff_one_iff_fderiv.1 (h.of_le hn)).2 /-- If a function is at least `C^1`, its bundled derivative (mapping `(x, v)` to `Df(x) v`) is continuous. -/ diff --git a/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean b/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean index c399175ddd00d..a403a0f90c9d8 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean @@ -28,8 +28,11 @@ section FiniteDimensional open Function Module +open scoped ContDiff + variable [CompleteSpace 𝕜] + /-- A family of continuous linear maps is `C^n` on `s` if all its applications are. -/ theorem contDiffOn_clm_apply {f : D → E →L[𝕜] F} {s : Set D} [FiniteDimensional 𝕜 E] : ContDiffOn 𝕜 n f s ↔ ∀ y, ContDiffOn 𝕜 n (fun x => f x y) s := by @@ -53,20 +56,22 @@ domain and codomain (`D` and `E`). This is not the case for `contDiff_succ_iff_f often requires an inconvenient need to generalize `F`, which results in universe issues (see the discussion in the section of `ContDiff.comp`). -This lemma avoids these universe issues, but only applies for finite dimensional `E`. -/ -theorem contDiff_succ_iff_fderiv_apply [FiniteDimensional 𝕜 D] {n : ℕ} {f : D → E} : - ContDiff 𝕜 (n + 1) f ↔ Differentiable 𝕜 f ∧ ∀ y, ContDiff 𝕜 n fun x => fderiv 𝕜 f x y := by +This lemma avoids these universe issues, but only applies for finite dimensional `D`. -/ +theorem contDiff_succ_iff_fderiv_apply [FiniteDimensional 𝕜 D] : + ContDiff 𝕜 (n + 1) f ↔ Differentiable 𝕜 f ∧ + (n = ω → AnalyticOnNhd 𝕜 f Set.univ) ∧ ∀ y, ContDiff 𝕜 n fun x => fderiv 𝕜 f x y := by rw [contDiff_succ_iff_fderiv, contDiff_clm_apply_iff] -theorem contDiffOn_succ_of_fderiv_apply [FiniteDimensional 𝕜 D] {n : ℕ} {f : D → E} {s : Set D} - (hf : DifferentiableOn 𝕜 f s) (h : ∀ y, ContDiffOn 𝕜 n (fun x => fderivWithin 𝕜 f s x y) s) : +theorem contDiffOn_succ_of_fderiv_apply [FiniteDimensional 𝕜 D] + (hf : DifferentiableOn 𝕜 f s) (h'f : n = ω → AnalyticOn 𝕜 f s) + (h : ∀ y, ContDiffOn 𝕜 n (fun x => fderivWithin 𝕜 f s x y) s) : ContDiffOn 𝕜 (n + 1) f s := - contDiffOn_succ_of_fderivWithin hf <| contDiffOn_clm_apply.mpr h + contDiffOn_succ_of_fderivWithin hf h'f <| contDiffOn_clm_apply.mpr h -theorem contDiffOn_succ_iff_fderiv_apply [FiniteDimensional 𝕜 D] {n : ℕ} {f : D → E} {s : Set D} - (hs : UniqueDiffOn 𝕜 s) : +theorem contDiffOn_succ_iff_fderiv_apply [FiniteDimensional 𝕜 D] (hs : UniqueDiffOn 𝕜 s) : ContDiffOn 𝕜 (n + 1) f s ↔ - DifferentiableOn 𝕜 f s ∧ ∀ y, ContDiffOn 𝕜 n (fun x => fderivWithin 𝕜 f s x y) s := by + DifferentiableOn 𝕜 f s ∧ (n = ω → AnalyticOn 𝕜 f s) ∧ + ∀ y, ContDiffOn 𝕜 n (fun x => fderivWithin 𝕜 f s x y) s := by rw [contDiffOn_succ_iff_fderivWithin hs, contDiffOn_clm_apply] end FiniteDimensional diff --git a/Mathlib/Analysis/Calculus/ContDiff/WithLp.lean b/Mathlib/Analysis/Calculus/ContDiff/WithLp.lean index 4c99b51e455ba..7c1f15c0a1dd3 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/WithLp.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/WithLp.lean @@ -17,24 +17,24 @@ open ContinuousLinearMap variable {𝕜 ι : Type*} {E : ι → Type*} {H : Type*} variable [NontriviallyNormedField 𝕜] [NormedAddCommGroup H] [∀ i, NormedAddCommGroup (E i)] [∀ i, NormedSpace 𝕜 (E i)] [NormedSpace 𝕜 H] [Fintype ι] (p) [Fact (1 ≤ p)] - {f : H → PiLp p E} {f' : H →L[𝕜] PiLp p E} {t : Set H} {y : H} + {n : WithTop ℕ∞} {f : H → PiLp p E} {f' : H →L[𝕜] PiLp p E} {t : Set H} {y : H} -theorem contDiffWithinAt_piLp {n : ℕ∞} : +theorem contDiffWithinAt_piLp : ContDiffWithinAt 𝕜 n f t y ↔ ∀ i, ContDiffWithinAt 𝕜 n (fun x => f x i) t y := by rw [← (PiLp.continuousLinearEquiv p 𝕜 E).comp_contDiffWithinAt_iff, contDiffWithinAt_pi] rfl -theorem contDiffAt_piLp {n : ℕ∞} : +theorem contDiffAt_piLp : ContDiffAt 𝕜 n f y ↔ ∀ i, ContDiffAt 𝕜 n (fun x => f x i) y := by rw [← (PiLp.continuousLinearEquiv p 𝕜 E).comp_contDiffAt_iff, contDiffAt_pi] rfl -theorem contDiffOn_piLp {n : ℕ∞} : +theorem contDiffOn_piLp : ContDiffOn 𝕜 n f t ↔ ∀ i, ContDiffOn 𝕜 n (fun x => f x i) t := by rw [← (PiLp.continuousLinearEquiv p 𝕜 E).comp_contDiffOn_iff, contDiffOn_pi] rfl -theorem contDiff_piLp {n : ℕ∞} : ContDiff 𝕜 n f ↔ ∀ i, ContDiff 𝕜 n fun x => f x i := by +theorem contDiff_piLp : ContDiff 𝕜 n f ↔ ∀ i, ContDiff 𝕜 n fun x => f x i := by rw [← (PiLp.continuousLinearEquiv p 𝕜 E).comp_contDiff_iff, contDiff_pi] rfl diff --git a/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean b/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean index b4f9b4db76e68..f57cd0a063cfd 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean @@ -95,7 +95,7 @@ lemma fderivWithin_fderivWithin_eq_of_mem_nhdsWithin (h : t ∈ 𝓝[s] x) fderivWithin 𝕜 (fderivWithin 𝕜 f s) s x = fderivWithin 𝕜 (fderivWithin 𝕜 f t) t x := by have A : ∀ᶠ y in 𝓝[s] x, fderivWithin 𝕜 f s y = fderivWithin 𝕜 f t y := by have : ∀ᶠ y in 𝓝[s] x, ContDiffWithinAt 𝕜 2 f t y := - nhdsWithin_le_iff.2 h (nhdsWithin_mono _ (subset_insert x t) hf.eventually) + nhdsWithin_le_iff.2 h (nhdsWithin_mono _ (subset_insert x t) (hf.eventually (by simp))) filter_upwards [self_mem_nhdsWithin, this, eventually_eventually_nhdsWithin.2 h] with y hy h'y h''y exact fderivWithin_of_mem_nhdsWithin h''y (hs y hy) (h'y.differentiableWithinAt one_le_two) @@ -412,11 +412,11 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] [IsRCLikeNormedField 𝕜 [NormedSpace 𝕜 F] {s : Set E} {f : E → F} {x : E} theorem second_derivative_symmetric_of_eventually {f' : E → E →L[𝕜] F} {x : E} - {f'' : E →L[𝕜] E →L[𝕜] F} (hf : ∀ᶠ y in 𝓝 x, HasFDerivAt f (f' y) y) (hx : HasFDerivAt f' f'' x) - (v w : E) : f'' v w = f'' w v := by - letI := IsRCLikeNormedField.rclike 𝕜 - letI : NormedSpace ℝ E := NormedSpace.restrictScalars ℝ 𝕜 E - letI : NormedSpace ℝ F := NormedSpace.restrictScalars ℝ 𝕜 F + {f'' : E →L[𝕜] E →L[𝕜] F} (hf : ∀ᶠ y in 𝓝 x, HasFDerivAt f (f' y) y) + (hx : HasFDerivAt f' f'' x) (v w : E) : f'' v w = f'' w v := by + let _ := IsRCLikeNormedField.rclike 𝕜 + let _ : NormedSpace ℝ E := NormedSpace.restrictScalars ℝ 𝕜 E + let _ : NormedSpace ℝ F := NormedSpace.restrictScalars ℝ 𝕜 F let f'R : E → E →L[ℝ] F := fun x ↦ (f' x).restrictScalars ℝ have hfR : ∀ᶠ y in 𝓝 x, HasFDerivAt f (f'R y) y := by filter_upwards [hf] with y hy using HasFDerivAt.restrictScalars ℝ hy @@ -448,7 +448,8 @@ theorem ContDiffAt.isSymmSndFDerivAt {n : WithTop ℕ∞} (hf : ContDiffAt 𝕜 IsSymmSndFDerivAt 𝕜 f x := by intro v w apply second_derivative_symmetric_of_eventually (f := f) (f' := fderiv 𝕜 f) (x := x) - · obtain ⟨u, hu, h'u⟩ : ∃ u ∈ 𝓝 x, ContDiffOn 𝕜 2 f u := hf.contDiffOn (m := 2) hn + · obtain ⟨u, hu, h'u⟩ : ∃ u ∈ 𝓝 x, ContDiffOn 𝕜 2 f u := + (hf.of_le hn).contDiffOn (m := 2) le_rfl (by simp) rcases mem_nhds_iff.1 hu with ⟨v, vu, v_open, xv⟩ filter_upwards [v_open.mem_nhds xv] with y hy have : DifferentiableAt 𝕜 f y := by @@ -468,7 +469,7 @@ theorem ContDiffWithinAt.isSymmSndFDerivWithinAt {n : WithTop ℕ∞} (hf : Cont /- We argue that, at interior points, the second derivative is symmetric, and moreover by continuity it converges to the second derivative at `x`. Therefore, the latter is also symmetric. -/ - rcases hf.contDiffOn' hn with ⟨u, u_open, xu, hu⟩ + rcases (hf.of_le hn).contDiffOn' le_rfl (by simp) with ⟨u, u_open, xu, hu⟩ simp only [insert_eq_of_mem h'x] at hu have h'u : UniqueDiffOn 𝕜 (s ∩ u) := hs.inter u_open obtain ⟨y, hy, y_lim⟩ : ∃ y, (∀ (n : ℕ), y n ∈ interior s) ∧ Tendsto y atTop (𝓝 x) := diff --git a/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean b/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean index 858642c83caa8..a013a881fc12f 100644 --- a/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean +++ b/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean @@ -133,12 +133,12 @@ theorem contDiffOn_of_differentiableOn_deriv {n : ℕ∞} continuous. -/ theorem ContDiffOn.continuousOn_iteratedDerivWithin {n : WithTop ℕ∞} {m : ℕ} (h : ContDiffOn 𝕜 n f s) - (hmn : (m : ℕ∞) ≤ n) (hs : UniqueDiffOn 𝕜 s) : ContinuousOn (iteratedDerivWithin m f s) s := by + (hmn : m ≤ n) (hs : UniqueDiffOn 𝕜 s) : ContinuousOn (iteratedDerivWithin m f s) s := by simpa only [iteratedDerivWithin_eq_equiv_comp, LinearIsometryEquiv.comp_continuousOn_iff] using h.continuousOn_iteratedFDerivWithin hmn hs theorem ContDiffWithinAt.differentiableWithinAt_iteratedDerivWithin {n : WithTop ℕ∞} {m : ℕ} - (h : ContDiffWithinAt 𝕜 n f s x) (hmn : (m : ℕ∞) < n) (hs : UniqueDiffOn 𝕜 (insert x s)) : + (h : ContDiffWithinAt 𝕜 n f s x) (hmn : m < n) (hs : UniqueDiffOn 𝕜 (insert x s)) : DifferentiableWithinAt 𝕜 (iteratedDerivWithin m f s) s x := by simpa only [iteratedDerivWithin_eq_equiv_comp, LinearIsometryEquiv.comp_differentiableWithinAt_iff] using @@ -159,6 +159,15 @@ theorem contDiffOn_iff_continuousOn_differentiableOn_deriv {n : ℕ∞} (hs : Un simp only [contDiffOn_iff_continuousOn_differentiableOn hs, iteratedFDerivWithin_eq_equiv_comp, LinearIsometryEquiv.comp_continuousOn_iff, LinearIsometryEquiv.comp_differentiableOn_iff] +/-- The property of being `C^n`, initially defined in terms of the Fréchet derivative, can be +reformulated in terms of the one-dimensional derivative on sets with unique derivatives. -/ +theorem contDiffOn_nat_iff_continuousOn_differentiableOn_deriv {n : ℕ} (hs : UniqueDiffOn 𝕜 s) : + ContDiffOn 𝕜 n f s ↔ (∀ m : ℕ, m ≤ n → ContinuousOn (iteratedDerivWithin m f s) s) ∧ + ∀ m : ℕ, m < n → DifferentiableOn 𝕜 (iteratedDerivWithin m f s) s := by + rw [show n = ((n : ℕ∞) : WithTop ℕ∞) from rfl, + contDiffOn_iff_continuousOn_differentiableOn_deriv hs] + simp + /-- The `n+1`-th iterated derivative within a set with unique derivatives can be obtained by differentiating the `n`-th iterated derivative. -/ theorem iteratedDerivWithin_succ {x : 𝕜} (hxs : UniqueDiffWithinAt 𝕜 s x) : @@ -230,6 +239,14 @@ theorem contDiff_iff_iteratedDeriv {n : ℕ∞} : ContDiff 𝕜 n f ↔ simp only [contDiff_iff_continuous_differentiable, iteratedFDeriv_eq_equiv_comp, LinearIsometryEquiv.comp_continuous_iff, LinearIsometryEquiv.comp_differentiable_iff] +/-- The property of being `C^n`, initially defined in terms of the Fréchet derivative, can be +reformulated in terms of the one-dimensional derivative. -/ +theorem contDiff_nat_iff_iteratedDeriv {n : ℕ} : ContDiff 𝕜 n f ↔ + (∀ m : ℕ, m ≤ n → Continuous (iteratedDeriv m f)) ∧ + ∀ m : ℕ, m < n → Differentiable 𝕜 (iteratedDeriv m f) := by + rw [show n = ((n : ℕ∞) : WithTop ℕ∞) from rfl, contDiff_iff_iteratedDeriv] + simp + /-- To check that a function is `n` times continuously differentiable, it suffices to check that its first `n` derivatives are differentiable. This is slightly too strong as the condition we require on the `n`-th derivative is differentiability instead of continuity, but it has the diff --git a/Mathlib/Analysis/Calculus/Taylor.lean b/Mathlib/Analysis/Calculus/Taylor.lean index 8a56fcf0d80a4..7d6b563e1c66a 100644 --- a/Mathlib/Analysis/Calculus/Taylor.lean +++ b/Mathlib/Analysis/Calculus/Taylor.lean @@ -119,8 +119,7 @@ theorem continuousOn_taylorWithinEval {f : ℝ → E} {x : ℝ} {n : ℕ} {s : S simp_rw [taylor_within_apply] refine continuousOn_finset_sum (Finset.range (n + 1)) fun i hi => ?_ refine (continuousOn_const.mul ((continuousOn_const.sub continuousOn_id).pow _)).smul ?_ - rw [show (n : WithTop ℕ∞) = (n : ℕ∞) by rfl, - contDiffOn_iff_continuousOn_differentiableOn_deriv hs] at hf + rw [contDiffOn_nat_iff_continuousOn_differentiableOn_deriv hs] at hf cases' hf with hf_left specialize hf_left i simp only [Finset.mem_range] at hi diff --git a/Mathlib/Analysis/Complex/CauchyIntegral.lean b/Mathlib/Analysis/Complex/CauchyIntegral.lean index 5bfbd617a7f48..340b252b8b626 100644 --- a/Mathlib/Analysis/Complex/CauchyIntegral.lean +++ b/Mathlib/Analysis/Complex/CauchyIntegral.lean @@ -577,9 +577,9 @@ theorem _root_.DifferentiableOn.analyticOn {s : Set ℂ} {f : ℂ → E} (hd : D /-- If `f : ℂ → E` is complex differentiable on some open set `s`, then it is continuously differentiable on `s`. -/ -protected theorem _root_.DifferentiableOn.contDiffOn {s : Set ℂ} {f : ℂ → E} {n : ℕ} +protected theorem _root_.DifferentiableOn.contDiffOn {s : Set ℂ} {f : ℂ → E} {n : WithTop ℕ∞} (hd : DifferentiableOn ℂ f s) (hs : IsOpen s) : ContDiffOn ℂ n f s := - (hd.analyticOnNhd hs).contDiffOn + (hd.analyticOnNhd hs).contDiffOn_of_completeSpace /-- A complex differentiable function `f : ℂ → E` is analytic at every point. -/ protected theorem _root_.Differentiable.analyticAt {f : ℂ → E} (hf : Differentiable ℂ f) (z : ℂ) : @@ -587,7 +587,8 @@ protected theorem _root_.Differentiable.analyticAt {f : ℂ → E} (hf : Differe hf.differentiableOn.analyticAt univ_mem /-- A complex differentiable function `f : ℂ → E` is continuously differentiable at every point. -/ -protected theorem _root_.Differentiable.contDiff {f : ℂ → E} (hf : Differentiable ℂ f) {n : ℕ∞} : +protected theorem _root_.Differentiable.contDiff + {f : ℂ → E} (hf : Differentiable ℂ f) {n : WithTop ℕ∞} : ContDiff ℂ n f := contDiff_iff_contDiffAt.mpr fun z ↦ (hf.analyticAt z).contDiffAt diff --git a/Mathlib/Analysis/Convolution.lean b/Mathlib/Analysis/Convolution.lean index f2e6fd1a550aa..9363b3cee07de 100644 --- a/Mathlib/Analysis/Convolution.lean +++ b/Mathlib/Analysis/Convolution.lean @@ -1190,7 +1190,7 @@ theorem contDiffOn_convolution_right_with_param_aux {G : Type uP} {E' : Type uP} HasFDerivAt (fun q : P × G => (f ⋆[L, μ] g q.1) q.2) (f' q₀.1 q₀.2) q₀ := hasFDerivAt_convolution_right_with_param L hs hk hgs hf hg.one_of_succ rw [contDiffOn_succ_iff_fderiv_of_isOpen (hs.prod (@isOpen_univ G _))] at hg ⊢ - refine ⟨?_, ?_⟩ + refine ⟨?_, by simp, ?_⟩ · rintro ⟨p, x⟩ ⟨hp, -⟩ exact (A (p, x) hp).differentiableAt.differentiableWithinAt · suffices H : ContDiffOn 𝕜 n (↿f') (s ×ˢ univ) by @@ -1207,9 +1207,9 @@ theorem contDiffOn_convolution_right_with_param_aux {G : Type uP} {E' : Type uP} rintro ⟨p, y⟩ ⟨hp, hy⟩ exact hgs p y hp hy apply ih (L.precompR (P × G) : _) B - convert hg.2 + convert hg.2.2 | htop ih => - rw [contDiffOn_top] at hg ⊢ + rw [contDiffOn_infty] at hg ⊢ exact fun n ↦ ih n L hgs (hg n) /-- The convolution `f * g` is `C^n` when `f` is locally integrable and `g` is `C^n` and compactly diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index b1f19e08a826b..40e9573a8eb34 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -545,7 +545,7 @@ lemma _root_.Function.HasTemperateGrowth.of_fderiv {f : E → F} (h'f : Function.HasTemperateGrowth (fderiv ℝ f)) (hf : Differentiable ℝ f) {k : ℕ} {C : ℝ} (h : ∀ x, ‖f x‖ ≤ C * (1 + ‖x‖) ^ k) : Function.HasTemperateGrowth f := by - refine ⟨contDiff_top_iff_fderiv.2 ⟨hf, h'f.1⟩ , fun n ↦ ?_⟩ + refine ⟨contDiff_succ_iff_fderiv.2 ⟨hf, by simp, h'f.1⟩ , fun n ↦ ?_⟩ rcases n with rfl|m · exact ⟨k, C, fun x ↦ by simpa using h x⟩ · rcases h'f.2 m with ⟨k', C', h'⟩ @@ -934,7 +934,7 @@ variable [RCLike 𝕜] [NormedSpace 𝕜 F] [SMulCommClass ℝ 𝕜 F] def fderivCLM : 𝓢(E, F) →L[𝕜] 𝓢(E, E →L[ℝ] F) := mkCLM (fderiv ℝ) (fun f g _ => fderiv_add f.differentiableAt g.differentiableAt) (fun a f _ => fderiv_const_smul f.differentiableAt a) - (fun f => (contDiff_top_iff_fderiv.mp f.smooth').2) fun ⟨k, n⟩ => + (fun f => (contDiff_succ_iff_fderiv.mp f.smooth').2.2) fun ⟨k, n⟩ => ⟨{⟨k, n + 1⟩}, 1, zero_le_one, fun f x => by simpa only [schwartzSeminormFamily_apply, Seminorm.comp_apply, Finset.sup_singleton, one_smul, norm_iteratedFDeriv_fderiv, one_mul] using f.le_seminorm 𝕜 k (n + 1) x⟩ @@ -947,7 +947,7 @@ theorem fderivCLM_apply (f : 𝓢(E, F)) (x : E) : fderivCLM 𝕜 f x = fderiv def derivCLM : 𝓢(ℝ, F) →L[𝕜] 𝓢(ℝ, F) := mkCLM deriv (fun f g _ => deriv_add f.differentiableAt g.differentiableAt) (fun a f _ => deriv_const_smul a f.differentiableAt) - (fun f => (contDiff_top_iff_deriv.mp f.smooth').2) fun ⟨k, n⟩ => + (fun f => (contDiff_succ_iff_deriv.mp f.smooth').2.2) fun ⟨k, n⟩ => ⟨{⟨k, n + 1⟩}, 1, zero_le_one, fun f x => by simpa only [Real.norm_eq_abs, Finset.sup_singleton, schwartzSeminormFamily_apply, one_mul, norm_iteratedFDeriv_eq_norm_iteratedDeriv, ← iteratedDeriv_succ'] using diff --git a/Mathlib/Analysis/InnerProductSpace/Calculus.lean b/Mathlib/Analysis/InnerProductSpace/Calculus.lean index 6fff4e5ede2f4..dcbeab2b02667 100644 --- a/Mathlib/Analysis/InnerProductSpace/Calculus.lean +++ b/Mathlib/Analysis/InnerProductSpace/Calculus.lean @@ -61,7 +61,7 @@ theorem differentiable_inner : Differentiable ℝ fun p : E × E => ⟪p.1, p.2 variable (𝕜) variable {G : Type*} [NormedAddCommGroup G] [NormedSpace ℝ G] {f g : G → E} {f' g' : G →L[ℝ] E} - {s : Set G} {x : G} {n : ℕ∞} + {s : Set G} {x : G} {n : WithTop ℕ∞} theorem ContDiffWithinAt.inner (hf : ContDiffWithinAt ℝ n f s x) (hg : ContDiffWithinAt ℝ n g s x) : ContDiffWithinAt ℝ n (fun x => ⟪f x, g x⟫) s x := @@ -302,19 +302,19 @@ theorem hasFDerivWithinAt_euclidean : ∀ i, HasFDerivWithinAt (fun x => f x i) (PiLp.proj _ _ i ∘L f') t y := hasFDerivWithinAt_piLp _ -theorem contDiffWithinAt_euclidean {n : ℕ∞} : +theorem contDiffWithinAt_euclidean {n : WithTop ℕ∞} : ContDiffWithinAt 𝕜 n f t y ↔ ∀ i, ContDiffWithinAt 𝕜 n (fun x => f x i) t y := contDiffWithinAt_piLp _ -theorem contDiffAt_euclidean {n : ℕ∞} : +theorem contDiffAt_euclidean {n : WithTop ℕ∞} : ContDiffAt 𝕜 n f y ↔ ∀ i, ContDiffAt 𝕜 n (fun x => f x i) y := contDiffAt_piLp _ -theorem contDiffOn_euclidean {n : ℕ∞} : +theorem contDiffOn_euclidean {n : WithTop ℕ∞} : ContDiffOn 𝕜 n f t ↔ ∀ i, ContDiffOn 𝕜 n (fun x => f x i) t := contDiffOn_piLp _ -theorem contDiff_euclidean {n : ℕ∞} : ContDiff 𝕜 n f ↔ ∀ i, ContDiff 𝕜 n fun x => f x i := +theorem contDiff_euclidean {n : WithTop ℕ∞} : ContDiff 𝕜 n f ↔ ∀ i, ContDiff 𝕜 n fun x => f x i := contDiff_piLp _ end PiLike diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean index 9ebb2a8c9c154..2e44424c0033c 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean @@ -672,8 +672,8 @@ def prodL : @[simps! apply symm_apply] def piₗᵢ {ι' : Type v'} [Fintype ι'] {E' : ι' → Type wE'} [∀ i', NormedAddCommGroup (E' i')] [∀ i', NormedSpace 𝕜 (E' i')] : - (∀ i', ContinuousMultilinearMap 𝕜 E (E' i')) ≃ₗᵢ[𝕜] - ContinuousMultilinearMap 𝕜 E (∀ i, E' i) where + (Π i', ContinuousMultilinearMap 𝕜 E (E' i')) + ≃ₗᵢ[𝕜] (ContinuousMultilinearMap 𝕜 E (Π i, E' i)) where toLinearEquiv := piLinearEquiv norm_map' := opNorm_pi diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/Analytic.lean b/Mathlib/Analysis/SpecialFunctions/Complex/Analytic.lean index d717114a27b82..ef1748b715e87 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/Analytic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/Analytic.lean @@ -62,11 +62,11 @@ theorem AnalyticAt.cpow (fa : AnalyticAt ℂ f x) (ga : AnalyticAt ℂ g x) exact fa.cpow ga m /-- `f z ^ g z` is analytic if `f z` avoids nonpositive reals -/ -theorem AnalyticOnNhd.cpow (fs : AnalyticOnNhd ℂ f s) (gs : AnalyticOnNhd ℂ g s) - (m : ∀ z ∈ s, f z ∈ slitPlane) : AnalyticOnNhd ℂ (fun z ↦ f z ^ g z) s := +theorem AnalyticOn.cpow (fs : AnalyticOn ℂ f s) (gs : AnalyticOn ℂ g s) + (m : ∀ z ∈ s, f z ∈ slitPlane) : AnalyticOn ℂ (fun z ↦ f z ^ g z) s := fun z n ↦ (fs z n).cpow (gs z n) (m z n) /-- `f z ^ g z` is analytic if `f z` avoids nonpositive reals -/ -theorem AnalyticOn.cpow (fs : AnalyticOn ℂ f s) (gs : AnalyticOn ℂ g s) - (m : ∀ z ∈ s, f z ∈ slitPlane) : AnalyticOn ℂ (fun z ↦ f z ^ g z) s := +theorem AnalyticOnNhd.cpow (fs : AnalyticOnNhd ℂ f s) (gs : AnalyticOnNhd ℂ g s) + (m : ∀ z ∈ s, f z ∈ slitPlane) : AnalyticOnNhd ℂ (fun z ↦ f z ^ g z) s := fun z n ↦ (fs z n).cpow (gs z n) (m z n) diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/LogDeriv.lean b/Mathlib/Analysis/SpecialFunctions/Complex/LogDeriv.lean index 464284b57e5ea..305f06d9e1523 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/LogDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/LogDeriv.lean @@ -64,7 +64,7 @@ theorem hasStrictFDerivAt_log_real {x : ℂ} (h : x ∈ slitPlane) : HasStrictFDerivAt log (x⁻¹ • (1 : ℂ →L[ℝ] ℂ)) x := (hasStrictDerivAt_log h).complexToReal_fderiv -theorem contDiffAt_log {x : ℂ} (h : x ∈ slitPlane) {n : ℕ∞} : ContDiffAt ℂ n log x := +theorem contDiffAt_log {x : ℂ} (h : x ∈ slitPlane) {n : WithTop ℕ∞} : ContDiffAt ℂ n log x := expPartialHomeomorph.contDiffAt_symm_deriv (exp_ne_zero <| log x) h (hasDerivAt_exp _) contDiff_exp.contDiffAt diff --git a/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean b/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean index 7a9ffe7ff6f72..805c933b00240 100644 --- a/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne -/ import Mathlib.Analysis.Complex.RealDeriv -import Mathlib.Analysis.Calculus.ContDiff.Analytic import Mathlib.Analysis.Calculus.ContDiff.RCLike import Mathlib.Analysis.Calculus.FDeriv.Analytic import Mathlib.Analysis.Calculus.IteratedDeriv.Lemmas diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean index efd1c239aa8d8..19a01e4add05d 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean @@ -383,6 +383,21 @@ theorem isLittleO_const_log_atTop {c : ℝ} : (fun _ => c) =o[atTop] log := by filter_upwards [eventually_gt_atTop 1] with x hx aesop (add safe forward log_pos) +/-- `Real.exp` as a `PartialHomeomorph` with `source = univ` and `target = {z | 0 < z}`. -/ +@[simps] noncomputable def expPartialHomeomorph : PartialHomeomorph ℝ ℝ where + toFun := Real.exp + invFun := Real.log + source := univ + target := Ioi (0 : ℝ) + map_source' x _ := exp_pos x + map_target' _ _ := mem_univ _ + left_inv' _ _ := by simp + right_inv' _ hx := exp_log hx + open_source := isOpen_univ + open_target := isOpen_Ioi + continuousOn_toFun := continuousOn_exp + continuousOn_invFun x hx := (continuousAt_log (ne_of_gt hx)).continuousWithinAt + end Real section Continuity diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean index b886b2c966010..f1eb252e9bb51 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean @@ -66,15 +66,27 @@ theorem deriv_log (x : ℝ) : deriv log x = x⁻¹ := theorem deriv_log' : deriv log = Inv.inv := funext deriv_log +theorem contDiffAt_log {n : WithTop ℕ∞} {x : ℝ} : ContDiffAt ℝ n log x ↔ x ≠ 0 := by + refine ⟨fun h ↦ continuousAt_log_iff.1 h.continuousAt, fun hx ↦ ?_⟩ + have A y (hy : 0 < y) : ContDiffAt ℝ n log y := by + apply expPartialHomeomorph.contDiffAt_symm_deriv (f₀' := y) hy.ne' (by simpa) + · convert hasDerivAt_exp (log y) + rw [exp_log hy] + · exact analyticAt_rexp.contDiffAt + rcases hx.lt_or_lt with hx | hx + · have : ContDiffAt ℝ n (log ∘ (fun y ↦ -y)) x := by + apply ContDiffAt.comp + apply A _ (Left.neg_pos_iff.mpr hx) + apply contDiffAt_id.neg + convert this + ext x + simp + · exact A x hx + theorem contDiffOn_log {n : WithTop ℕ∞} : ContDiffOn ℝ n log {0}ᶜ := by - suffices ContDiffOn ℝ ω log {0}ᶜ from this.of_le le_top - rw [← contDiffOn_infty_iff_contDiffOn_omega] - refine (contDiffOn_top_iff_deriv_of_isOpen isOpen_compl_singleton).2 ?_ - simp [differentiableOn_log, contDiffOn_inv] - -theorem contDiffAt_log {n : WithTop ℕ∞} : ContDiffAt ℝ n log x ↔ x ≠ 0 := - ⟨fun h => continuousAt_log_iff.1 h.continuousAt, fun hx => - (contDiffOn_log x hx).contDiffAt <| IsOpen.mem_nhds isOpen_compl_singleton hx⟩ + intro x hx + simp only [mem_compl_iff, mem_singleton_iff] at hx + exact (contDiffAt_log.2 hx).contDiffWithinAt end Real diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean index 0c619bd07f671..39bf892635542 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean @@ -368,8 +368,9 @@ theorem contDiff_rpow_const_of_le {p : ℝ} {n : ℕ} (h : ↑n ≤ p) : · exact contDiff_zero.2 (continuous_id.rpow_const fun x => Or.inr <| by simpa using h) · have h1 : 1 ≤ p := le_trans (by simp) h rw [Nat.cast_succ, ← le_sub_iff_add_le] at h - rw [show ((n + 1 : ℕ) : WithTop ℕ∞) = n + 1 from rfl, contDiff_succ_iff_deriv, - deriv_rpow_const' h1] + rw [show ((n + 1 : ℕ) : WithTop ℕ∞) = n + 1 from rfl, + contDiff_succ_iff_deriv, deriv_rpow_const' h1] + simp only [WithTop.natCast_ne_top, analyticOn_univ, IsEmpty.forall_iff, true_and] exact ⟨differentiable_rpow_const h1, contDiff_const.mul (ihn h)⟩ theorem contDiffAt_rpow_const_of_le {x p : ℝ} {n : ℕ} (h : ↑n ≤ p) : @@ -394,7 +395,7 @@ open Real section fderiv variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] {f g : E → ℝ} {f' g' : E →L[ℝ] ℝ} - {x : E} {s : Set E} {c p : ℝ} {n : ℕ∞} + {x : E} {s : Set E} {c p : ℝ} {n : WithTop ℕ∞} #adaptation_note /-- https://github.com/leanprover/lean4/pull/6024 added `by exact` to deal with unification issues. -/ diff --git a/Mathlib/Analysis/SpecialFunctions/SmoothTransition.lean b/Mathlib/Analysis/SpecialFunctions/SmoothTransition.lean index a9815ec827dfa..c66c0166fed3f 100644 --- a/Mathlib/Analysis/SpecialFunctions/SmoothTransition.lean +++ b/Mathlib/Analysis/SpecialFunctions/SmoothTransition.lean @@ -108,7 +108,8 @@ theorem contDiff_polynomial_eval_inv_mul {n : ℕ∞} (p : ℝ[X]) : induction m generalizing p with | zero => exact contDiff_zero.2 <| continuous_polynomial_eval_inv_mul _ | succ m ihm => - refine contDiff_succ_iff_deriv.2 ⟨differentiable_polynomial_eval_inv_mul _, ?_⟩ + rw [show ((m + 1 : ℕ) : WithTop ℕ∞) = m + 1 from rfl] + refine contDiff_succ_iff_deriv.2 ⟨differentiable_polynomial_eval_inv_mul _, by simp, ?_⟩ convert ihm (X ^ 2 * (p - derivative (R := ℝ) p)) using 2 exact (hasDerivAt_polynomial_eval_inv_mul p _).deriv diff --git a/Mathlib/Analysis/SpecialFunctions/Sqrt.lean b/Mathlib/Analysis/SpecialFunctions/Sqrt.lean index 4a26cd60c9cbc..aefa91cf42110 100644 --- a/Mathlib/Analysis/SpecialFunctions/Sqrt.lean +++ b/Mathlib/Analysis/SpecialFunctions/Sqrt.lean @@ -57,7 +57,7 @@ theorem deriv_sqrt_aux {x : ℝ} (hx : x ≠ 0) : theorem hasStrictDerivAt_sqrt {x : ℝ} (hx : x ≠ 0) : HasStrictDerivAt (√·) (1 / (2 * √x)) x := (deriv_sqrt_aux hx).1 -theorem contDiffAt_sqrt {x : ℝ} {n : ℕ∞} (hx : x ≠ 0) : ContDiffAt ℝ n (√·) x := +theorem contDiffAt_sqrt {x : ℝ} {n : WithTop ℕ∞} (hx : x ≠ 0) : ContDiffAt ℝ n (√·) x := (deriv_sqrt_aux hx).2 n theorem hasDerivAt_sqrt {x : ℝ} (hx : x ≠ 0) : HasDerivAt (√·) (1 / (2 * √x)) x := @@ -98,8 +98,8 @@ end deriv section fderiv -variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] {f : E → ℝ} {n : ℕ∞} {s : Set E} - {x : E} {f' : E →L[ℝ] ℝ} +variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] {f : E → ℝ} {n : WithTop ℕ∞} + {s : Set E} {x : E} {f' : E →L[ℝ] ℝ} theorem HasFDerivAt.sqrt (hf : HasFDerivAt f f' x) (hx : f x ≠ 0) : HasFDerivAt (fun y => √(f y)) ((1 / (2 * √(f x))) • f') x := diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/ArctanDeriv.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/ArctanDeriv.lean index 43102430e6996..106f49d5a9a1d 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/ArctanDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/ArctanDeriv.lean @@ -55,7 +55,7 @@ theorem deriv_tan (x : ℝ) : deriv tan x = 1 / cos x ^ 2 := else (hasDerivAt_tan h).deriv @[simp] -theorem contDiffAt_tan {n x} : ContDiffAt ℝ n tan x ↔ cos x ≠ 0 := +theorem contDiffAt_tan {n : WithTop ℕ∞} {x : ℝ} : ContDiffAt ℝ n tan x ↔ cos x ≠ 0 := ⟨fun h => continuousAt_tan.1 h.continuousAt, fun h => (Complex.contDiffAt_tan.2 <| mod_cast h).real_of_complex⟩ @@ -88,7 +88,7 @@ theorem differentiable_arctan : Differentiable ℝ arctan := theorem deriv_arctan : deriv arctan = fun (x : ℝ) => 1 / (1 + x ^ 2) := funext fun x => (hasDerivAt_arctan x).deriv -theorem contDiff_arctan {n : ℕ∞} : ContDiff ℝ n arctan := +theorem contDiff_arctan {n : WithTop ℕ∞} : ContDiff ℝ n arctan := contDiff_iff_contDiffAt.2 fun x => have : cos (arctan x) ≠ 0 := (cos_arctan_pos x).ne' tanPartialHomeomorph.contDiffAt_symm_deriv (by simpa) trivial (hasDerivAt_tan this) diff --git a/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean b/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean index c9dd661d44afd..cae9a80c72889 100644 --- a/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean +++ b/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean @@ -204,12 +204,11 @@ class SmoothInv₀ {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [To /-- Inversion is smooth away from `0`. -/ smoothAt_inv₀ : ∀ ⦃x : G⦄, x ≠ 0 → ContMDiffAt I I ⊤ (fun y ↦ y⁻¹) x -instance {𝕜 : Type*} [NontriviallyNormedField 𝕜] [CompleteSpace 𝕜] : SmoothInv₀ 𝓘(𝕜) 𝕜 := - { smoothAt_inv₀ := by - intro x hx - change ContMDiffAt 𝓘(𝕜) 𝓘(𝕜) ⊤ Inv.inv x - rw [contMDiffAt_iff_contDiffAt] - exact contDiffAt_inv 𝕜 hx } +instance {𝕜 : Type*} [NontriviallyNormedField 𝕜] : SmoothInv₀ 𝓘(𝕜) 𝕜 where + smoothAt_inv₀ x hx := by + change ContMDiffAt 𝓘(𝕜) 𝓘(𝕜) ⊤ Inv.inv x + rw [contMDiffAt_iff_contDiffAt] + exact contDiffAt_inv 𝕜 hx variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [TopologicalSpace H] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] (I : ModelWithCorners 𝕜 E H) {G : Type*} diff --git a/Mathlib/Geometry/Manifold/AnalyticManifold.lean b/Mathlib/Geometry/Manifold/AnalyticManifold.lean index ea26292a58a0d..cecf164328810 100644 --- a/Mathlib/Geometry/Manifold/AnalyticManifold.lean +++ b/Mathlib/Geometry/Manifold/AnalyticManifold.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Lee, Geoffrey Irving -/ import Mathlib.Analysis.Analytic.Constructions -import Mathlib.Analysis.Calculus.ContDiff.Analytic import Mathlib.Geometry.Manifold.SmoothManifoldWithCorners /-! @@ -15,9 +14,6 @@ half-space (to get manifolds with boundaries) for which changes of coordinates a interior and smooth everywhere (including at the boundary). The definition mirrors `SmoothManifoldWithCorners`, but using an `analyticGroupoid` in place of `contDiffGroupoid`. All analytic manifolds are smooth manifolds. - -Completeness is required throughout, but this is nonessential: it is due to many of the lemmas about -AnalyticOn` requiring completeness for ease of proof. -/ noncomputable section @@ -167,11 +163,10 @@ instance AnalyticManifold.prod {E A : Type} [NormedAddCommGroup E] [NormedSpace /-- Analytic manifolds are smooth manifolds. -/ instance AnalyticManifold.smoothManifoldWithCorners [ChartedSpace H M] - [cm : AnalyticManifold I M] [CompleteSpace E] : + [cm : AnalyticManifold I M] : SmoothManifoldWithCorners I M where - compatible := by - intro f g hf hg - have m := cm.compatible hf hg - exact ⟨m.1.contDiffOn, m.2.contDiffOn⟩ + compatible hf hg := ⟨(cm.compatible hf hg).1.contDiffOn I.uniqueDiffOn_preimage_source, + (cm.compatible hg hf).1.contDiffOn I.uniqueDiffOn_preimage_source⟩ + end AnalyticManifold diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean index 49996ee7ba8cb..112fab89cc177 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean @@ -563,8 +563,8 @@ theorem ContMDiff.continuous (hf : ContMDiff I I' n f) : Continuous f := theorem contMDiffWithinAt_top : ContMDiffWithinAt I I' ⊤ f s x ↔ ∀ n : ℕ, ContMDiffWithinAt I I' n f s x := - ⟨fun h n => ⟨h.1, contDiffWithinAt_top.1 h.2 n⟩, fun H => - ⟨(H 0).1, contDiffWithinAt_top.2 fun n => (H n).2⟩⟩ + ⟨fun h n => ⟨h.1, contDiffWithinAt_infty.1 h.2 n⟩, fun H => + ⟨(H 0).1, contDiffWithinAt_infty.2 fun n => (H n).2⟩⟩ theorem contMDiffAt_top : ContMDiffAt I I' ⊤ f x ↔ ∀ n : ℕ, ContMDiffAt I I' n f x := contMDiffWithinAt_top @@ -690,7 +690,7 @@ theorem contMDiffWithinAt_iff_contMDiffOn_nhds (hu _ (mem_of_mem_nhdsWithin hxs hmem)).mono_of_mem_nhdsWithin hmem⟩ -- The property is true in charts. Let `v` be a good neighborhood in the chart where the function -- is smooth. - rcases (contMDiffWithinAt_iff'.1 h).2.contDiffOn le_rfl with ⟨v, hmem, hsub, hv⟩ + rcases (contMDiffWithinAt_iff'.1 h).2.contDiffOn le_rfl (by simp) with ⟨v, hmem, hsub, hv⟩ have hxs' : extChartAt I x x ∈ (extChartAt I x).target ∩ (extChartAt I x).symm ⁻¹' (s ∩ f ⁻¹' (extChartAt I' (f x)).source) := ⟨(extChartAt I x).map_source (mem_extChartAt_source _), by rwa [extChartAt_to_inv], by diff --git a/Mathlib/Geometry/Manifold/Sheaf/LocallyRingedSpace.lean b/Mathlib/Geometry/Manifold/Sheaf/LocallyRingedSpace.lean index 5c91ede6dc259..787dce0f8a583 100644 --- a/Mathlib/Geometry/Manifold/Sheaf/LocallyRingedSpace.lean +++ b/Mathlib/Geometry/Manifold/Sheaf/LocallyRingedSpace.lean @@ -35,7 +35,7 @@ universe u Later, replace with `open scoped ContDiff`. -/ local notation "∞" => (⊤ : ℕ∞) -variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] [CompleteSpace 𝕜] +variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] {EM : Type*} [NormedAddCommGroup EM] [NormedSpace 𝕜 EM] {HM : Type*} [TopologicalSpace HM] (IM : ModelWithCorners 𝕜 EM HM) {M : Type u} [TopologicalSpace M] [ChartedSpace HM M] diff --git a/Mathlib/Tactic/FunProp/ContDiff.lean b/Mathlib/Tactic/FunProp/ContDiff.lean index 037d1812296a5..b5325b5527b64 100644 --- a/Mathlib/Tactic/FunProp/ContDiff.lean +++ b/Mathlib/Tactic/FunProp/ContDiff.lean @@ -61,7 +61,7 @@ variable {K : Type*} [NontriviallyNormedField K] variable {E : Type*} [NormedAddCommGroup E] [NormedSpace K E] variable {s} -theorem ContDiffOn.div' [CompleteSpace K] {f g : E → K} {n} (hf : ContDiffOn K n f s) +theorem ContDiffOn.div' {f g : E → K} {n} (hf : ContDiffOn K n f s) (hg : ContDiffOn K n g s) (h₀ : ∀ x ∈ s, g x ≠ 0) : ContDiffOn K n (fun x => f x / g x) s := ContDiffOn.div hf hg h₀ From cee87aa25cfbd8e4c3cd89bc26cce86e927e36dc Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Fri, 29 Nov 2024 22:48:10 +0000 Subject: [PATCH 421/829] refactor(AlgebraicGeometry/EllipticCurve/*): add `WeierstrassCurve.IsElliptic` and remove `EllipticCurve` (#18531) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The original `EllipticCurve` is `WeierstrassCurve` bundled with `Δ' : Rˣ` and a proof that `Δ' = Δ`. This definition makes `Δ'` and `j` computable, but at the expense that some APIs must be duplicated for `EllipticCurve` and `WeierstrassCurve`. This PR adds the typeclass `WeierstrassCurve.IsElliptic` which asserts that `IsUnit W.Δ`, and which is a `Prop`. The old API for elliptic curves are replaced by a `WeierstrassCurve` `W` together with `[W.IsElliptic]`. The caveat is that `Δ'` and `j` becomes noncomputable. The original design of this PR is `WeierstrassCurve.Elliptic` which is `Invertible W.Δ`, and which is a `Type`. By this design `Δ'` and `j` remain computable. But this design makes simpNF linter timeout. So this design is abandoned and replaced by current design. Also rename `J.lean` -> `IsomOfJ.lean` and split `ModelsWithJ.lean`. --- Mathlib.lean | 3 +- .../EllipticCurve/Affine.lean | 19 +- .../EllipticCurve/Group.lean | 10 +- .../EllipticCurve/{J.lean => IsomOfJ.lean} | 43 ++- .../EllipticCurve/ModelsWithJ.lean | 196 +++++++++++ .../EllipticCurve/NormalForms.lean | 121 ++----- .../EllipticCurve/VariableChange.lean | 106 +++--- .../EllipticCurve/Weierstrass.lean | 317 ++++-------------- 8 files changed, 375 insertions(+), 440 deletions(-) rename Mathlib/AlgebraicGeometry/EllipticCurve/{J.lean => IsomOfJ.lean} (91%) create mode 100644 Mathlib/AlgebraicGeometry/EllipticCurve/ModelsWithJ.lean diff --git a/Mathlib.lean b/Mathlib.lean index 5ce336f55f28b..027c728331798 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -943,8 +943,9 @@ import Mathlib.AlgebraicGeometry.EllipticCurve.Affine import Mathlib.AlgebraicGeometry.EllipticCurve.DivisionPolynomial.Basic import Mathlib.AlgebraicGeometry.EllipticCurve.DivisionPolynomial.Degree import Mathlib.AlgebraicGeometry.EllipticCurve.Group -import Mathlib.AlgebraicGeometry.EllipticCurve.J +import Mathlib.AlgebraicGeometry.EllipticCurve.IsomOfJ import Mathlib.AlgebraicGeometry.EllipticCurve.Jacobian +import Mathlib.AlgebraicGeometry.EllipticCurve.ModelsWithJ import Mathlib.AlgebraicGeometry.EllipticCurve.NormalForms import Mathlib.AlgebraicGeometry.EllipticCurve.Projective import Mathlib.AlgebraicGeometry.EllipticCurve.VariableChange diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean index 231b6e92006e9..189a1bc957349 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean @@ -14,7 +14,7 @@ This file defines the type of points on a Weierstrass curve as an inductive, con at infinity and affine points satisfying a Weierstrass equation with a nonsingular condition. This file also defines the negation and addition operations of the group law for this type, and proves that they respect the Weierstrass equation and the nonsingular condition. The fact that they form an -abelian group is proven in `Mathlib.AlgebraicGeometry.EllipticCurve.Group`. +abelian group is proven in `Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean`. ## Mathematical background @@ -65,7 +65,7 @@ The group law on this set is then uniquely determined by these constructions. * `WeierstrassCurve.Affine.nonsingular_add`: addition preserves the nonsingular condition. * `WeierstrassCurve.Affine.nonsingular_of_Δ_ne_zero`: an affine Weierstrass curve is nonsingular at every point if its discriminant is non-zero. - * `EllipticCurve.Affine.nonsingular`: an affine elliptic curve is nonsingular at every point. + * `WeierstrassCurve.Affine.nonsingular`: an affine elliptic curve is nonsingular at every point. ## Notations @@ -948,21 +948,16 @@ end BaseChange @[deprecated (since := "2024-06-03")] alias baseChange_addY' := baseChange_negAddY @[deprecated (since := "2024-06-03")] alias map_addY' := map_negAddY -end WeierstrassCurve.Affine - /-! ## Elliptic curves -/ -/-- The coercion from an elliptic curve to a Weierstrass curve in affine coordinates. -/ -abbrev EllipticCurve.toAffine {R : Type u} [CommRing R] (E : EllipticCurve R) : - WeierstrassCurve.Affine R := - E.toWeierstrassCurve.toAffine - -namespace EllipticCurve.Affine +section EllipticCurve -variable {R : Type u} [CommRing R] (E : EllipticCurve R) +variable {R : Type u} [CommRing R] (E : WeierstrassCurve R) [E.IsElliptic] lemma nonsingular [Nontrivial R] {x y : R} (h : E.toAffine.Equation x y) : E.toAffine.Nonsingular x y := E.toAffine.nonsingular_of_Δ_ne_zero h <| E.coe_Δ' ▸ E.Δ'.ne_zero -end EllipticCurve.Affine +end EllipticCurve + +end WeierstrassCurve.Affine diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean index c01a58c13679a..251acb43cbb94 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean @@ -12,8 +12,8 @@ import Mathlib.RingTheory.Polynomial.UniqueFactorization # Group law on Weierstrass curves This file proves that the nonsingular rational points on a Weierstrass curve forms an abelian group -under the geometric group law defined in `Mathlib.AlgebraicGeometry.EllipticCurve.Affine` and in -`Mathlib.AlgebraicGeometry.EllipticCurve.Jacobian`. +under the geometric group law defined in `Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean` and +in `Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean`. ## Mathematical background @@ -582,14 +582,14 @@ noncomputable instance : AddCommGroup W.Point where end WeierstrassCurve.Jacobian.Point -namespace EllipticCurve.Affine.Point +namespace WeierstrassCurve.Affine.Point /-! ## Elliptic curves in affine coordinates -/ -variable {R : Type} [Nontrivial R] [CommRing R] (E : EllipticCurve R) +variable {R : Type*} [Nontrivial R] [CommRing R] (E : WeierstrassCurve R) [E.IsElliptic] /-- An affine point on an elliptic curve `E` over `R`. -/ def mk {x y : R} (h : E.toAffine.Equation x y) : E.toAffine.Point := WeierstrassCurve.Affine.Point.some <| nonsingular E h -end EllipticCurve.Affine.Point +end WeierstrassCurve.Affine.Point diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/J.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/IsomOfJ.lean similarity index 91% rename from Mathlib/AlgebraicGeometry/EllipticCurve/J.lean rename to Mathlib/AlgebraicGeometry/EllipticCurve/IsomOfJ.lean index 3a9bbae95d05d..3474c6e51b8e4 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/J.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/IsomOfJ.lean @@ -8,33 +8,32 @@ import Mathlib.FieldTheory.IsSepClosed /-! -# Further properties of j-invariants of elliptic curves - -This file states some further properties of j-invariants of elliptic curves. +# Elliptic curves with same j-invariants are isomorphic ## Main results -- `EllipticCurve.exists_variableChange_of_j_eq`: if `E` and `E'` are elliptic curves with the same - `j`-invariants defined over a separably closed field, then there exists a change of variables +- `WeierstrassCurve.exists_variableChange_of_j_eq`: if `E` and `E'` are elliptic curves with the + same `j`-invariants defined over a separably closed field, then there exists a change of variables over that field which change `E` into `E'`. -/ -open WeierstrassCurve Polynomial +open Polynomial variable {F : Type*} [Field F] [IsSepClosed F] -namespace EllipticCurve +namespace WeierstrassCurve -variable (E E' : EllipticCurve F) +variable (E E' : WeierstrassCurve F) [E.IsElliptic] [E'.IsElliptic] section CharTwo variable [CharP F 2] +omit [E.IsElliptic] [E'.IsElliptic] in private lemma exists_variableChange_of_char_two_of_j_ne_zero [E.IsCharTwoJNeZeroNF] [E'.IsCharTwoJNeZeroNF] (heq : E.a₆ = E'.a₆) : - ∃ C : WeierstrassCurve.VariableChange F, E.variableChange C = E' := by + ∃ C : VariableChange F, E.variableChange C = E' := by obtain ⟨s, hs⟩ := IsSepClosed.exists_root_C_mul_X_pow_add_C_mul_X_add_C' 2 2 1 1 (E.a₂ + E'.a₂) (by norm_num) (by norm_num) one_ne_zero use ⟨1, 0, s, 0⟩ @@ -53,7 +52,7 @@ private lemma exists_variableChange_of_char_two_of_j_ne_zero private lemma exists_variableChange_of_char_two_of_j_eq_zero [E.IsCharTwoJEqZeroNF] [E'.IsCharTwoJEqZeroNF] : - ∃ C : WeierstrassCurve.VariableChange F, E.variableChange C = E' := by + ∃ C : VariableChange F, E.variableChange C = E' := by have ha₃ := E.Δ'.ne_zero rw [E.coe_Δ', Δ_of_isCharTwoJEqZeroNF_of_char_two, pow_ne_zero_iff (Nat.succ_ne_zero _)] at ha₃ have ha₃' := E'.Δ'.ne_zero @@ -86,7 +85,7 @@ private lemma exists_variableChange_of_char_two_of_j_eq_zero linear_combination ht - (t ^ 2 + E.a₃ * t) * CharP.cast_eq_zero F 2 private lemma exists_variableChange_of_char_two (heq : E.j = E'.j) : - ∃ C : WeierstrassCurve.VariableChange F, E.variableChange C = E' := by + ∃ C : VariableChange F, E.variableChange C = E' := by obtain ⟨C, _ | _⟩ := E.exists_variableChange_isCharTwoNF · obtain ⟨C', _ | _⟩ := E'.exists_variableChange_isCharTwoNF · simp_rw [← variableChange_j E C, ← variableChange_j E' C', @@ -94,7 +93,7 @@ private lemma exists_variableChange_of_char_two (heq : E.j = E'.j) : obtain ⟨C'', hC⟩ := exists_variableChange_of_char_two_of_j_ne_zero _ _ heq use (C'.inv.comp C'').comp C rw [variableChange_comp, variableChange_comp, hC, ← variableChange_comp, - WeierstrassCurve.VariableChange.comp_left_inv, variableChange_id] + VariableChange.comp_left_inv, variableChange_id] · have h := (E.variableChange C).j_ne_zero_of_isCharTwoJNeZeroNF_of_char_two rw [variableChange_j, heq, ← variableChange_j E' C', j_of_isCharTwoJEqZeroNF_of_char_two] at h @@ -108,7 +107,7 @@ private lemma exists_variableChange_of_char_two (heq : E.j = E'.j) : (E.variableChange C) (E'.variableChange C') use (C'.inv.comp C'').comp C rw [variableChange_comp, variableChange_comp, hC, ← variableChange_comp, - WeierstrassCurve.VariableChange.comp_left_inv, variableChange_id] + VariableChange.comp_left_inv, variableChange_id] end CharTwo @@ -118,7 +117,7 @@ variable [CharP F 3] private lemma exists_variableChange_of_char_three_of_j_ne_zero [E.IsCharThreeJNeZeroNF] [E'.IsCharThreeJNeZeroNF] (heq : E.j = E'.j) : - ∃ C : WeierstrassCurve.VariableChange F, E.variableChange C = E' := by + ∃ C : VariableChange F, E.variableChange C = E' := by have h := E.Δ'.ne_zero rw [E.coe_Δ', Δ_of_isCharThreeJNeZeroNF_of_char_three, mul_ne_zero_iff, neg_ne_zero, pow_ne_zero_iff three_ne_zero] at h @@ -155,7 +154,7 @@ private lemma exists_variableChange_of_char_three_of_j_ne_zero private lemma exists_variableChange_of_char_three_of_j_eq_zero [E.IsShortNF] [E'.IsShortNF] : - ∃ C : WeierstrassCurve.VariableChange F, E.variableChange C = E' := by + ∃ C : VariableChange F, E.variableChange C = E' := by have ha₄ := E.Δ'.ne_zero rw [E.coe_Δ', Δ_of_isShortNF_of_char_three, neg_ne_zero, pow_ne_zero_iff three_ne_zero] at ha₄ have ha₄' := E'.Δ'.ne_zero @@ -188,14 +187,14 @@ private lemma exists_variableChange_of_char_three_of_j_eq_zero linear_combination hr private lemma exists_variableChange_of_char_three (heq : E.j = E'.j) : - ∃ C : WeierstrassCurve.VariableChange F, E.variableChange C = E' := by + ∃ C : VariableChange F, E.variableChange C = E' := by obtain ⟨C, _ | _⟩ := E.exists_variableChange_isCharThreeNF · obtain ⟨C', _ | _⟩ := E'.exists_variableChange_isCharThreeNF · rw [← variableChange_j E C, ← variableChange_j E' C'] at heq obtain ⟨C'', hC⟩ := exists_variableChange_of_char_three_of_j_ne_zero _ _ heq use (C'.inv.comp C'').comp C rw [variableChange_comp, variableChange_comp, hC, ← variableChange_comp, - WeierstrassCurve.VariableChange.comp_left_inv, variableChange_id] + VariableChange.comp_left_inv, variableChange_id] · have h := (E.variableChange C).j_ne_zero_of_isCharThreeJNeZeroNF_of_char_three rw [variableChange_j, heq, ← variableChange_j E' C', j_of_isShortNF_of_char_three] at h exact False.elim (h rfl) @@ -207,7 +206,7 @@ private lemma exists_variableChange_of_char_three (heq : E.j = E'.j) : (E.variableChange C) (E'.variableChange C') use (C'.inv.comp C'').comp C rw [variableChange_comp, variableChange_comp, hC, ← variableChange_comp, - WeierstrassCurve.VariableChange.comp_left_inv, variableChange_id] + VariableChange.comp_left_inv, variableChange_id] end CharThree @@ -215,7 +214,7 @@ section CharNeTwoOrThree private lemma exists_variableChange_of_char_ne_two_or_three {p : ℕ} [CharP F p] (hchar2 : p ≠ 2) (hchar3 : p ≠ 3) (heq : E.j = E'.j) : - ∃ C : WeierstrassCurve.VariableChange F, E.variableChange C = E' := by + ∃ C : VariableChange F, E.variableChange C = E' := by replace hchar2 : (2 : F) ≠ 0 := CharP.cast_ne_zero_of_ne_of_prime F Nat.prime_two hchar2 replace hchar3 : (3 : F) ≠ 0 := CharP.cast_ne_zero_of_ne_of_prime F Nat.prime_three hchar3 haveI := NeZero.mk hchar2 @@ -239,7 +238,7 @@ private lemma exists_variableChange_of_char_ne_two_or_three rw [← variableChange_j E' C] at heq obtain ⟨C', hC⟩ := this _ heq hE' exact ⟨C.inv.comp C', by rw [variableChange_comp, hC, ← variableChange_comp, - WeierstrassCurve.VariableChange.comp_left_inv, variableChange_id]⟩ + VariableChange.comp_left_inv, variableChange_id]⟩ simp_rw [j, Units.val_inv_eq_inv_val, inv_mul_eq_div, div_eq_div_iff E.Δ'.ne_zero E'.Δ'.ne_zero, coe_Δ', Δ_of_isShortNF, c₄_of_isShortNF] at heq replace heq : E.a₄ ^ 3 * E'.a₆ ^ 2 = E'.a₄ ^ 3 * E.a₆ ^ 2 := by @@ -337,7 +336,7 @@ end CharNeTwoOrThree separably closed field, then there exists a change of variables over that field which change one curve into another. -/ theorem exists_variableChange_of_j_eq (heq : E.j = E'.j) : - ∃ C : WeierstrassCurve.VariableChange F, E.variableChange C = E' := by + ∃ C : VariableChange F, E.variableChange C = E' := by obtain ⟨p, _⟩ := CharP.exists F by_cases hchar2 : p = 2 · subst hchar2 @@ -347,4 +346,4 @@ theorem exists_variableChange_of_j_eq (heq : E.j = E'.j) : exact exists_variableChange_of_char_three _ _ heq exact exists_variableChange_of_char_ne_two_or_three _ _ hchar2 hchar3 heq -end EllipticCurve +end WeierstrassCurve diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/ModelsWithJ.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/ModelsWithJ.lean new file mode 100644 index 0000000000000..31e52e81f9fef --- /dev/null +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/ModelsWithJ.lean @@ -0,0 +1,196 @@ +/- +Copyright (c) 2021 Kevin Buzzard. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kevin Buzzard, David Kurniadi Angdinata +-/ +import Mathlib.AlgebraicGeometry.EllipticCurve.Weierstrass + +/-! +# Models of elliptic curves with prescribed j-invariant + +This file defines the Weierstrass equation over a field with prescribed j-invariant, +proved that it is an elliptic curve, and that its j-invariant is equal to the given value. +It is a modification of [silverman2009], Chapter III, Proposition 1.4 (c). + +## Main definitions + + * `WeierstrassCurve.ofJ0`: an elliptic curve whose j-invariant is 0. + * `WeierstrassCurve.ofJ1728`: an elliptic curve whose j-invariant is 1728. + * `WeierstrassCurve.ofJNe0Or1728`: an elliptic curve whose j-invariant is neither 0 nor 1728. + * `WeierstrassCurve.ofJ`: an elliptic curve whose j-invariant equal to j. + +## Main statements + + * `WeierstrassCurve.ofJ_j`: the j-invariant of `WeierstrassCurve.ofJ` is equal to j. + +## References + + * [J Silverman, *The Arithmetic of Elliptic Curves*][silverman2009] + +## Tags + +elliptic curve, weierstrass equation, j invariant +-/ + +namespace WeierstrassCurve + +variable (R : Type*) [CommRing R] (W : WeierstrassCurve R) + +/-- The Weierstrass curve $Y^2 + Y = X^3$. It is of j-invariant 0 if it is an elliptic curve. -/ +def ofJ0 : WeierstrassCurve R := + ⟨0, 0, 1, 0, 0⟩ + +lemma ofJ0_c₄ : (ofJ0 R).c₄ = 0 := by + rw [ofJ0, c₄, b₂, b₄] + norm_num1 + +lemma ofJ0_Δ : (ofJ0 R).Δ = -27 := by + rw [ofJ0, Δ, b₂, b₄, b₆, b₈] + norm_num1 + +/-- The Weierstrass curve $Y^2 = X^3 + X$. It is of j-invariant 1728 if it is an elliptic curve. -/ +def ofJ1728 : WeierstrassCurve R := + ⟨0, 0, 0, 1, 0⟩ + +lemma ofJ1728_c₄ : (ofJ1728 R).c₄ = -48 := by + rw [ofJ1728, c₄, b₂, b₄] + norm_num1 + +lemma ofJ1728_Δ : (ofJ1728 R).Δ = -64 := by + rw [ofJ1728, Δ, b₂, b₄, b₆, b₈] + norm_num1 + +variable {R} (j : R) + +/-- The Weierstrass curve $Y^2 + (j - 1728)XY = X^3 - 36(j - 1728)^3X - (j - 1728)^5$. +It is a modification of the curve in [silverman2009], Chapter III, Proposition 1.4 (c) to avoid +denominators. It is of j-invariant j if it is an elliptic curve. -/ +def ofJNe0Or1728 : WeierstrassCurve R := + ⟨j - 1728, 0, 0, -36 * (j - 1728) ^ 3, -(j - 1728) ^ 5⟩ + +lemma ofJNe0Or1728_c₄ : (ofJNe0Or1728 j).c₄ = j * (j - 1728) ^ 3 := by + simp only [ofJNe0Or1728, c₄, b₂, b₄] + ring1 + +lemma ofJNe0Or1728_Δ : (ofJNe0Or1728 j).Δ = j ^ 2 * (j - 1728) ^ 9 := by + simp only [ofJNe0Or1728, Δ, b₂, b₄, b₆, b₈] + ring1 + +variable (R) [W.IsElliptic] + +-- TODO: change to `[IsUnit ...]` once #17458 is merged +/-- When 3 is a unit, $Y^2 + Y = X^3$ is an elliptic curve. +It is of j-invariant 0 (see `WeierstrassCurve.ofJ0_j`). -/ +instance [hu : Fact (IsUnit (3 : R))] : (ofJ0 R).IsElliptic := by + rw [isElliptic_iff, ofJ0_Δ] + convert (hu.out.pow 3).neg + norm_num1 + +-- TODO: change to `[IsUnit ...]` once #17458 is merged +lemma ofJ0_j [Fact (IsUnit (3 : R))] : (ofJ0 R).j = 0 := by + rw [j, ofJ0_c₄] + ring1 + +-- TODO: change to `[IsUnit ...]` once #17458 is merged +/-- When 2 is a unit, $Y^2 = X^3 + X$ is an elliptic curve. +It is of j-invariant 1728 (see `WeierstrassCurve.ofJ1728_j`). -/ +instance [hu : Fact (IsUnit (2 : R))] : (ofJ1728 R).IsElliptic := by + rw [isElliptic_iff, ofJ1728_Δ] + convert (hu.out.pow 6).neg + norm_num1 + +-- TODO: change to `[IsUnit ...]` once #17458 is merged +lemma ofJ1728_j [Fact (IsUnit (2 : R))] : (ofJ1728 R).j = 1728 := by + rw [j, Units.inv_mul_eq_iff_eq_mul, ofJ1728_c₄, coe_Δ', ofJ1728_Δ] + norm_num1 + +variable {R} + +-- TODO: change to `[IsUnit ...]` once #17458 is merged +/-- When j and j - 1728 are both units, +$Y^2 + (j - 1728)XY = X^3 - 36(j - 1728)^3X - (j - 1728)^5$ is an elliptic curve. +It is of j-invariant j (see `WeierstrassCurve.ofJNe0Or1728_j`). -/ +instance (j : R) [h1 : Fact (IsUnit j)] [h2 : Fact (IsUnit (j - 1728))] : + (ofJNe0Or1728 j).IsElliptic := by + rw [isElliptic_iff, ofJNe0Or1728_Δ] + exact (h1.out.pow 2).mul (h2.out.pow 9) + +-- TODO: change to `[IsUnit ...]` once #17458 is merged +lemma ofJNe0Or1728_j (j : R) [Fact (IsUnit j)] [Fact (IsUnit (j - 1728))] : + (ofJNe0Or1728 j).j = j := by + rw [WeierstrassCurve.j, Units.inv_mul_eq_iff_eq_mul, ofJNe0Or1728_c₄, coe_Δ', ofJNe0Or1728_Δ] + ring1 + +variable {F : Type*} [Field F] (j : F) [DecidableEq F] + +/-- For any element j of a field $F$, there exists an elliptic curve over $F$ +with j-invariant equal to j (see `WeierstrassCurve.ofJ_j`). +Its coefficients are given explicitly (see `WeierstrassCurve.ofJ0`, `WeierstrassCurve.ofJ1728` +and `WeierstrassCurve.ofJNe0Or1728`). -/ +def ofJ : WeierstrassCurve F := + if j = 0 then if (3 : F) = 0 then ofJ1728 F else ofJ0 F + else if j = 1728 then ofJ1728 F else ofJNe0Or1728 j + +lemma ofJ_0_of_three_ne_zero (h3 : (3 : F) ≠ 0) : ofJ 0 = ofJ0 F := by + rw [ofJ, if_pos rfl, if_neg h3] + +lemma ofJ_0_of_three_eq_zero (h3 : (3 : F) = 0) : ofJ 0 = ofJ1728 F := by + rw [ofJ, if_pos rfl, if_pos h3] + +lemma ofJ_0_of_two_eq_zero (h2 : (2 : F) = 0) : ofJ 0 = ofJ0 F := by + rw [ofJ, if_pos rfl, if_neg ((show (3 : F) = 1 by linear_combination h2) ▸ one_ne_zero)] + +lemma ofJ_1728_of_three_eq_zero (h3 : (3 : F) = 0) : ofJ 1728 = ofJ1728 F := by + rw [ofJ, if_pos (by linear_combination 576 * h3), if_pos h3] + +lemma ofJ_1728_of_two_ne_zero (h2 : (2 : F) ≠ 0) : ofJ 1728 = ofJ1728 F := by + by_cases h3 : (3 : F) = 0 + · exact ofJ_1728_of_three_eq_zero h3 + · rw [ofJ, show (1728 : F) = 2 ^ 6 * 3 ^ 3 by norm_num1, + if_neg (mul_ne_zero (pow_ne_zero 6 h2) (pow_ne_zero 3 h3)), if_pos rfl] + +lemma ofJ_1728_of_two_eq_zero (h2 : (2 : F) = 0) : ofJ 1728 = ofJ0 F := by + rw [ofJ, if_pos (by linear_combination 864 * h2), + if_neg ((show (3 : F) = 1 by linear_combination h2) ▸ one_ne_zero)] + +lemma ofJ_ne_0_ne_1728 (h0 : j ≠ 0) (h1728 : j ≠ 1728) : ofJ j = ofJNe0Or1728 j := by + rw [ofJ, if_neg h0, if_neg h1728] + +instance : (ofJ j).IsElliptic := by + by_cases h0 : j = 0 + · by_cases h3 : (3 : F) = 0 + · have := Fact.mk (isUnit_of_mul_eq_one (2 : F) 2 (by linear_combination h3)) + rw [h0, ofJ_0_of_three_eq_zero h3] + infer_instance + · have := Fact.mk (Ne.isUnit h3) + rw [h0, ofJ_0_of_three_ne_zero h3] + infer_instance + · by_cases h1728 : j = 1728 + · have h2 : (2 : F) ≠ 0 := fun h ↦ h0 (by linear_combination h1728 + 864 * h) + have := Fact.mk h2.isUnit + rw [h1728, ofJ_1728_of_two_ne_zero h2] + infer_instance + · have := Fact.mk (Ne.isUnit h0) + have := Fact.mk (sub_ne_zero.2 h1728).isUnit + rw [ofJ_ne_0_ne_1728 j h0 h1728] + infer_instance + +lemma ofJ_j : (ofJ j).j = j := by + by_cases h0 : j = 0 + · by_cases h3 : (3 : F) = 0 + · have := Fact.mk (isUnit_of_mul_eq_one (2 : F) 2 (by linear_combination h3)) + simp_rw [h0, ofJ_0_of_three_eq_zero h3, ofJ1728_j] + linear_combination 576 * h3 + · have := Fact.mk (Ne.isUnit h3) + simp_rw [h0, ofJ_0_of_three_ne_zero h3, ofJ0_j] + · by_cases h1728 : j = 1728 + · have h2 : (2 : F) ≠ 0 := fun h ↦ h0 (by linear_combination h1728 + 864 * h) + have := Fact.mk h2.isUnit + simp_rw [h1728, ofJ_1728_of_two_ne_zero h2, ofJ1728_j] + · have := Fact.mk (Ne.isUnit h0) + have := Fact.mk (sub_ne_zero.2 h1728).isUnit + simp_rw [ofJ_ne_0_ne_1728 j h0 h1728, ofJNe0Or1728_j] + +instance : Inhabited { W : WeierstrassCurve F // W.IsElliptic } := ⟨⟨ofJ 37, inferInstance⟩⟩ + +end WeierstrassCurve diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/NormalForms.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/NormalForms.lean index 23a9a69744985..5738f216555e7 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/NormalForms.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/NormalForms.lean @@ -92,8 +92,7 @@ elliptic curve, weierstrass equation, normal form -/ -variable {R : Type*} [CommRing R] (W : WeierstrassCurve R) -variable {F : Type*} [Field F] (E : EllipticCurve F) +variable {R : Type*} [CommRing R] {F : Type*} [Field F] (W : WeierstrassCurve R) namespace WeierstrassCurve @@ -156,7 +155,7 @@ end Quantity section VariableChange -variable (E : EllipticCurve R) [Invertible (2 : R)] +variable [Invertible (2 : R)] /-- There is an explicit change of variables of a `WeierstrassCurve` to a normal form of characteristic ≠ 2, provided that 2 is invertible in the ring. -/ @@ -170,14 +169,6 @@ theorem exists_variableChange_isCharNeTwoNF : ∃ C : VariableChange R, (W.variableChange C).IsCharNeTwoNF := ⟨_, W.toCharNeTwoNF_spec⟩ -instance _root_.EllipticCurve.toCharNeTwoNF_spec : - (E.variableChange E.toCharNeTwoNF).IsCharNeTwoNF := - E.toWeierstrassCurve.toCharNeTwoNF_spec - -theorem _root_.EllipticCurve.exists_variableChange_isCharNeTwoNF : - ∃ C : VariableChange R, (E.variableChange C).IsCharNeTwoNF := - ⟨_, E.toCharNeTwoNF_spec⟩ - end VariableChange /-! ### Short normal form -/ @@ -248,25 +239,24 @@ theorem Δ_of_isShortNF_of_char_three : W.Δ = -W.a₄ ^ 3 := by rw [Δ_of_isShortNF] linear_combination (-21 * W.a₄ ^ 3 - 144 * W.a₆ ^ 2) * CharP.cast_eq_zero R 3 -variable [E.IsShortNF] +variable (W : WeierstrassCurve F) [W.IsElliptic] [W.IsShortNF] -theorem _root_.EllipticCurve.j_of_isShortNF : - E.j = 6912 * E.a₄ ^ 3 / (4 * E.a₄ ^ 3 + 27 * E.a₆ ^ 2) := by - have h := E.Δ'.ne_zero - rw [E.coe_Δ', Δ_of_isShortNF] at h - rw [EllipticCurve.j, Units.val_inv_eq_inv_val, ← div_eq_inv_mul, E.coe_Δ', +theorem j_of_isShortNF : W.j = 6912 * W.a₄ ^ 3 / (4 * W.a₄ ^ 3 + 27 * W.a₆ ^ 2) := by + have h := W.Δ'.ne_zero + rw [coe_Δ', Δ_of_isShortNF] at h + rw [j, Units.val_inv_eq_inv_val, ← div_eq_inv_mul, coe_Δ', c₄_of_isShortNF, Δ_of_isShortNF, div_eq_div_iff h (right_ne_zero_of_mul h)] ring1 @[simp] -theorem _root_.EllipticCurve.j_of_isShortNF_of_char_three [CharP F 3] : E.j = 0 := by - rw [EllipticCurve.j, c₄_of_isShortNF_of_char_three]; simp +theorem j_of_isShortNF_of_char_three [CharP F 3] : W.j = 0 := by + rw [j, c₄_of_isShortNF_of_char_three]; simp end Quantity section VariableChange -variable (E : EllipticCurve R) [Invertible (2 : R)] [Invertible (3 : R)] +variable [Invertible (2 : R)] [Invertible (3 : R)] /-- There is an explicit change of variables of a `WeierstrassCurve` to a short normal form, provided that 2 and 3 are invertible in the ring. @@ -282,13 +272,6 @@ theorem exists_variableChange_isShortNF : ∃ C : VariableChange R, (W.variableChange C).IsShortNF := ⟨_, W.toShortNF_spec⟩ -instance _root_.EllipticCurve.toShortNF_spec : (E.variableChange E.toShortNF).IsShortNF := - E.toWeierstrassCurve.toShortNF_spec - -theorem _root_.EllipticCurve.exists_variableChange_isShortNF : - ∃ C : VariableChange R, (E.variableChange C).IsShortNF := - ⟨_, E.toShortNF_spec⟩ - end VariableChange /-! ### Normal forms of characteristic = 3 and j ≠ 0 -/ @@ -360,21 +343,21 @@ theorem Δ_of_isCharThreeJNeZeroNF_of_char_three : W.Δ = -W.a₂ ^ 3 * W.a₆ : rw [Δ_of_isCharThreeJNeZeroNF] linear_combination (-21 * W.a₂ ^ 3 * W.a₆ - 144 * W.a₆ ^ 2) * CharP.cast_eq_zero R 3 -variable [E.IsCharThreeJNeZeroNF] [CharP F 3] +variable (W : WeierstrassCurve F) [W.IsElliptic] [W.IsCharThreeJNeZeroNF] [CharP F 3] @[simp] -theorem _root_.EllipticCurve.j_of_isCharThreeJNeZeroNF_of_char_three : E.j = -E.a₂ ^ 3 / E.a₆ := by - have h := E.Δ'.ne_zero - rw [E.coe_Δ', Δ_of_isCharThreeJNeZeroNF_of_char_three] at h - rw [EllipticCurve.j, Units.val_inv_eq_inv_val, ← div_eq_inv_mul, E.coe_Δ', +theorem j_of_isCharThreeJNeZeroNF_of_char_three : W.j = -W.a₂ ^ 3 / W.a₆ := by + have h := W.Δ'.ne_zero + rw [coe_Δ', Δ_of_isCharThreeJNeZeroNF_of_char_three] at h + rw [j, Units.val_inv_eq_inv_val, ← div_eq_inv_mul, coe_Δ', c₄_of_isCharThreeJNeZeroNF_of_char_three, Δ_of_isCharThreeJNeZeroNF_of_char_three, div_eq_div_iff h (right_ne_zero_of_mul h)] ring1 -theorem _root_.EllipticCurve.j_ne_zero_of_isCharThreeJNeZeroNF_of_char_three : E.j ≠ 0 := by - rw [E.j_of_isCharThreeJNeZeroNF_of_char_three, div_ne_zero_iff] - have h := E.Δ'.ne_zero - rwa [E.coe_Δ', Δ_of_isCharThreeJNeZeroNF_of_char_three, mul_ne_zero_iff] at h +theorem j_ne_zero_of_isCharThreeJNeZeroNF_of_char_three : W.j ≠ 0 := by + rw [j_of_isCharThreeJNeZeroNF_of_char_three, div_ne_zero_iff] + have h := W.Δ'.ne_zero + rwa [coe_Δ', Δ_of_isCharThreeJNeZeroNF_of_char_three, mul_ne_zero_iff] at h end Quantity @@ -420,10 +403,6 @@ theorem toShortNFOfCharThree_spec (hb₂ : W.b₂ = 0) : have H := W.toCharNeTwoNF_spec exact ⟨H.a₁, hb₂ ▸ W.toShortNFOfCharThree_a₂, H.a₃⟩ -theorem _root_.EllipticCurve.toShortNFOfCharThree_spec (E : EllipticCurve R) (hb₂ : E.b₂ = 0) : - (E.variableChange E.toShortNFOfCharThree).IsShortNF := - E.toWeierstrassCurve.toShortNFOfCharThree_spec hb₂ - variable (W : WeierstrassCurve F) /-- For a `WeierstrassCurve` defined over a field of characteristic = 3, @@ -455,14 +434,6 @@ theorem toCharThreeNF_spec_of_b₂_eq_zero (hb₂ : W.b₂ = 0) : VariableChange.id_comp] exact W.toShortNFOfCharThree_spec hb₂ -theorem _root_.EllipticCurve.toCharThreeNF_spec_of_b₂_ne_zero (hb₂ : E.b₂ ≠ 0) : - (E.variableChange E.toCharThreeNF).IsCharThreeJNeZeroNF := - E.toWeierstrassCurve.toCharThreeNF_spec_of_b₂_ne_zero hb₂ - -theorem _root_.EllipticCurve.toCharThreeNF_spec_of_b₂_eq_zero (hb₂ : E.b₂ = 0) : - (E.variableChange E.toCharThreeNF).IsShortNF := - E.toWeierstrassCurve.toCharThreeNF_spec_of_b₂_eq_zero hb₂ - instance toCharThreeNF_spec : (W.variableChange W.toCharThreeNF).IsCharThreeNF := by by_cases hb₂ : W.b₂ = 0 · haveI := W.toCharThreeNF_spec_of_b₂_eq_zero hb₂ @@ -474,14 +445,6 @@ theorem exists_variableChange_isCharThreeNF : ∃ C : VariableChange F, (W.variableChange C).IsCharThreeNF := ⟨_, W.toCharThreeNF_spec⟩ -instance _root_.EllipticCurve.toCharThreeNF_spec : - (E.variableChange E.toCharThreeNF).IsCharThreeNF := - E.toWeierstrassCurve.toCharThreeNF_spec - -theorem _root_.EllipticCurve.exists_variableChange_isCharThreeNF : - ∃ C : VariableChange F, (E.variableChange C).IsCharThreeNF := - ⟨_, E.toCharThreeNF_spec⟩ - end VariableChange /-! ### Normal forms of characteristic = 2 and j ≠ 0 -/ @@ -565,17 +528,17 @@ theorem Δ_of_isCharTwoJNeZeroNF_of_char_two : W.Δ = W.a₆ := by b₆_of_isCharTwoJNeZeroNF_of_char_two, b₈_of_isCharTwoJNeZeroNF_of_char_two] linear_combination -W.a₆ * CharP.cast_eq_zero R 2 -variable [E.IsCharTwoJNeZeroNF] [CharP F 2] +variable (W : WeierstrassCurve F) [W.IsElliptic] [W.IsCharTwoJNeZeroNF] [CharP F 2] @[simp] -theorem _root_.EllipticCurve.j_of_isCharTwoJNeZeroNF_of_char_two : E.j = 1 / E.a₆ := by - rw [EllipticCurve.j, Units.val_inv_eq_inv_val, ← div_eq_inv_mul, E.coe_Δ', +theorem j_of_isCharTwoJNeZeroNF_of_char_two : W.j = 1 / W.a₆ := by + rw [j, Units.val_inv_eq_inv_val, ← div_eq_inv_mul, coe_Δ', c₄_of_isCharTwoJNeZeroNF_of_char_two, Δ_of_isCharTwoJNeZeroNF_of_char_two, one_pow] -theorem _root_.EllipticCurve.j_ne_zero_of_isCharTwoJNeZeroNF_of_char_two : E.j ≠ 0 := by - rw [E.j_of_isCharTwoJNeZeroNF_of_char_two, div_ne_zero_iff] - have h := E.Δ'.ne_zero - rw [E.coe_Δ', Δ_of_isCharTwoJNeZeroNF_of_char_two] at h +theorem j_ne_zero_of_isCharTwoJNeZeroNF_of_char_two : W.j ≠ 0 := by + rw [j_of_isCharTwoJNeZeroNF_of_char_two, div_ne_zero_iff] + have h := W.Δ'.ne_zero + rw [coe_Δ', Δ_of_isCharTwoJNeZeroNF_of_char_two] at h exact ⟨one_ne_zero, h⟩ end Quantity @@ -651,19 +614,18 @@ theorem Δ_of_isCharTwoJEqZeroNF_of_char_two : W.Δ = W.a₃ ^ 4 := by rw [Δ_of_isCharTwoJEqZeroNF, b₆_of_char_two] linear_combination (-32 * W.a₄ ^ 3 - 14 * W.a₃ ^ 4) * CharP.cast_eq_zero R 2 -variable [E.IsCharTwoJEqZeroNF] +variable (W : WeierstrassCurve F) [W.IsElliptic] [W.IsCharTwoJEqZeroNF] -theorem _root_.EllipticCurve.j_of_isCharTwoJEqZeroNF : - E.j = 110592 * E.a₄ ^ 3 / (64 * E.a₄ ^ 3 + 27 * E.b₆ ^ 2) := by - have h := E.Δ'.ne_zero - rw [E.coe_Δ', Δ_of_isCharTwoJEqZeroNF] at h - rw [EllipticCurve.j, Units.val_inv_eq_inv_val, ← div_eq_inv_mul, E.coe_Δ', +theorem j_of_isCharTwoJEqZeroNF : W.j = 110592 * W.a₄ ^ 3 / (64 * W.a₄ ^ 3 + 27 * W.b₆ ^ 2) := by + have h := W.Δ'.ne_zero + rw [coe_Δ', Δ_of_isCharTwoJEqZeroNF] at h + rw [j, Units.val_inv_eq_inv_val, ← div_eq_inv_mul, coe_Δ', c₄_of_isCharTwoJEqZeroNF, Δ_of_isCharTwoJEqZeroNF, div_eq_div_iff h (neg_ne_zero.1 h)] ring1 @[simp] -theorem _root_.EllipticCurve.j_of_isCharTwoJEqZeroNF_of_char_two [CharP F 2] : E.j = 0 := by - rw [EllipticCurve.j, c₄_of_isCharTwoJEqZeroNF_of_char_two]; simp +theorem j_of_isCharTwoJEqZeroNF_of_char_two [CharP F 2] : W.j = 0 := by + rw [j, c₄_of_isCharTwoJEqZeroNF_of_char_two]; simp end Quantity @@ -698,10 +660,6 @@ theorem toCharTwoJEqZeroNF_spec (ha₁ : W.a₁ = 0) : · simp_rw [toCharTwoJEqZeroNF, variableChange_a₂, inv_one, Units.val_one] linear_combination 2 * W.a₂ * CharP.cast_eq_zero R 2 -theorem _root_.EllipticCurve.toCharTwoJEqZeroNF_spec (E : EllipticCurve R) (ha₁ : E.a₁ = 0) : - (E.variableChange E.toCharTwoJEqZeroNF).IsCharTwoJEqZeroNF := - E.toWeierstrassCurve.toCharTwoJEqZeroNF_spec ha₁ - variable (W : WeierstrassCurve F) /-- For a `WeierstrassCurve` defined over a field of characteristic = 2, @@ -719,10 +677,6 @@ theorem toCharTwoJNeZeroNF_spec (ha₁ : W.a₁ ≠ 0) : · field_simp [toCharTwoJNeZeroNF] linear_combination (W.a₁ ^ 4 * W.a₃ ^ 2 + W.a₁ ^ 5 * W.a₃ * W.a₂) * CharP.cast_eq_zero F 2 -theorem _root_.EllipticCurve.toCharTwoJNeZeroNF_spec (ha₁ : E.a₁ ≠ 0) : - (E.variableChange (E.toCharTwoJNeZeroNF ha₁)).IsCharTwoJNeZeroNF := - E.toWeierstrassCurve.toCharTwoJNeZeroNF_spec ha₁ - /-- For a `WeierstrassCurve` defined over a field of characteristic = 2, there is an explicit change of variables of it to `WeierstrassCurve.IsCharTwoNF`, that is, $Y^2 + XY = X^3 + a_2X^2 + a_6$ (`WeierstrassCurve.IsCharTwoJNeZeroNF`) or @@ -744,15 +698,6 @@ theorem exists_variableChange_isCharTwoNF : classical exact ⟨_, W.toCharTwoNF_spec⟩ -instance _root_.EllipticCurve.toCharTwoNF_spec [DecidableEq F] : - (E.variableChange E.toCharTwoNF).IsCharTwoNF := - E.toWeierstrassCurve.toCharTwoNF_spec - -theorem _root_.EllipticCurve.exists_variableChange_isCharTwoNF : - ∃ C : VariableChange F, (E.variableChange C).IsCharTwoNF := by - classical - exact ⟨_, E.toCharTwoNF_spec⟩ - end VariableChange end WeierstrassCurve diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/VariableChange.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/VariableChange.lean index 77ea4968688e9..5ee45845ec6d1 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/VariableChange.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/VariableChange.lean @@ -13,14 +13,13 @@ This file defines admissible linear change of variables of Weierstrass curves. ## Main definitions * `WeierstrassCurve.VariableChange`: a change of variables of Weierstrass curves. + * `WeierstrassCurve.VariableChange.instGroup`: change of variables form a group. * `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 + * `WeierstrassCurve.variableChange_j`: the j-invariant of an elliptic curve is invariant under an admissible linear change of variables. ## References @@ -47,7 +46,9 @@ section VariableChange /-- 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}$. -/ +$\begin{pmatrix} u^2 & 0 & r \cr u^2s & u^3 & t \cr 0 & 0 & 1 \end{pmatrix}$. +In other words, this is the change of variables $(X, Y) \mapsto (u^2X + r, u^3Y + u^2sX + t)$. +When `R` is a field, any two isomorphic Weierstrass equations are related by this. -/ @[ext] structure VariableChange (R : Type u) [CommRing R] where /-- The `u` coefficient of an admissible linear change of variables, which must be a unit. -/ @@ -91,9 +92,9 @@ 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 + · linear_combination -C.r * pow_mul_pow_eq_one 2 C.u.inv_mul + · linear_combination -C.s * C.u.inv_mul + · linear_combination (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 @@ -132,21 +133,21 @@ 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) + · linear_combination ↑C.u⁻¹ * C.s * 2 * C'.u.inv_mul + · linear_combination 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) + · linear_combination 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) + · linear_combination 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) + · linear_combination 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₃) @@ -201,6 +202,34 @@ lemma variableChange_Δ : (W.variableChange C).Δ = C.u⁻¹ ^ 12 * W.Δ := by variableChange_a₄, variableChange_a₆] ring1 +variable [W.IsElliptic] + +instance : (W.variableChange C).IsElliptic := by + rw [isElliptic_iff, variableChange_Δ] + exact (C.u⁻¹.isUnit.pow 12).mul W.isUnit_Δ + +set_option linter.docPrime false in +@[simp] +lemma variableChange_Δ' : (W.variableChange C).Δ' = C.u⁻¹ ^ 12 * W.Δ' := by + simp_rw [Units.ext_iff, Units.val_mul, coe_Δ', variableChange_Δ, Units.val_pow_eq_pow_val] + +set_option linter.docPrime false in +lemma coe_variableChange_Δ' : ((W.variableChange C).Δ' : R) = C.u⁻¹ ^ 12 * W.Δ' := by + simp_rw [coe_Δ', variableChange_Δ] + +set_option linter.docPrime false in +lemma inv_variableChange_Δ' : (W.variableChange C).Δ'⁻¹ = C.u ^ 12 * W.Δ'⁻¹ := by + rw [variableChange_Δ', mul_inv, inv_pow, inv_inv] + +set_option linter.docPrime false in +lemma coe_inv_variableChange_Δ' : (↑(W.variableChange C).Δ'⁻¹ : R) = C.u ^ 12 * W.Δ'⁻¹ := by + rw [inv_variableChange_Δ', Units.val_mul, Units.val_pow_eq_pow_val] + +@[simp] +lemma variableChange_j : (W.variableChange C).j = W.j := by + rw [j, coe_inv_variableChange_Δ', variableChange_c₄, j, mul_pow, ← pow_mul, ← mul_assoc, + mul_right_comm (C.u.val ^ 12), ← mul_pow, C.u.mul_inv, one_pow, one_mul] + end VariableChange section BaseChange @@ -270,56 +299,3 @@ lemma map_variableChange (C : VariableChange R) : 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 dcdc6fa4499d6..7265241ac9573 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean @@ -32,23 +32,15 @@ splitting field of `R` are precisely the $X$-coordinates of the non-zero 2-torsi * `WeierstrassCurve`: a Weierstrass curve over a commutative ring. * `WeierstrassCurve.Δ`: the discriminant of a Weierstrass curve. - * `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.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. - * `EllipticCurve.j`: the j-invariant of an elliptic curve. - * `EllipticCurve.ofJ0`: an elliptic curve whose j-invariant is 0. - * `EllipticCurve.ofJ1728`: an elliptic curve whose j-invariant is 1728. - * `EllipticCurve.ofJ'`: an elliptic curve whose j-invariant is neither 0 nor 1728. - * `EllipticCurve.ofJ`: an elliptic curve whose j-invariant equal to j. + * `WeierstrassCurve.IsElliptic`: typeclass asserting that a Weierstrass curve is an elliptic curve. + * `WeierstrassCurve.j`: the j-invariant of an elliptic curve. ## Main statements * `WeierstrassCurve.twoTorsionPolynomial_disc`: the discriminant of a Weierstrass curve is a constant factor of the cubic discriminant of its 2-torsion polynomial. - * `EllipticCurve.ofJ_j`: the j-invariant of `EllipticCurve.ofJ` is equal to j. ## Implementation notes @@ -345,122 +337,75 @@ lemma twoTorsionPolynomial_disc_of_char_three : W.twoTorsionPolynomial.disc = W. end CharThree -lemma twoTorsionPolynomial_disc_isUnit [Invertible (2 : R)] : +-- TODO: change to `[IsUnit ...]` once #17458 is merged +lemma twoTorsionPolynomial_disc_isUnit (hu : IsUnit (2 : R)) : IsUnit W.twoTorsionPolynomial.disc ↔ IsUnit W.Δ := by rw [twoTorsionPolynomial_disc, IsUnit.mul_iff, show (16 : R) = 2 ^ 4 by norm_num1] - exact and_iff_right <| isUnit_of_invertible <| 2 ^ 4 + exact and_iff_right <| hu.pow 4 -lemma twoTorsionPolynomial_disc_ne_zero [Nontrivial R] [Invertible (2 : R)] (hΔ : IsUnit W.Δ) : +-- TODO: change to `[IsUnit ...]` once #17458 is merged +-- TODO: In this case `IsUnit W.Δ` is just `W.IsElliptic`, consider removing/rephrasing this result +lemma twoTorsionPolynomial_disc_ne_zero [Nontrivial R] (hu : IsUnit (2 : R)) (hΔ : IsUnit W.Δ) : W.twoTorsionPolynomial.disc ≠ 0 := - (W.twoTorsionPolynomial_disc_isUnit.mpr hΔ).ne_zero + ((W.twoTorsionPolynomial_disc_isUnit hu).mpr hΔ).ne_zero end TorsionPolynomial -section ModelsWithJ - -/-! ### Models with prescribed j-invariant -/ - -variable (R) - -/-- The Weierstrass curve $Y^2 + Y = X^3$. It is of j-invariant 0 if it is an elliptic curve. -/ -def ofJ0 : WeierstrassCurve R := - ⟨0, 0, 1, 0, 0⟩ - -lemma ofJ0_c₄ : (ofJ0 R).c₄ = 0 := by - rw [ofJ0, c₄, b₂, b₄] - norm_num1 - -lemma ofJ0_Δ : (ofJ0 R).Δ = -27 := by - rw [ofJ0, Δ, b₂, b₄, b₆, b₈] - norm_num1 - -/-- The Weierstrass curve $Y^2 = X^3 + X$. It is of j-invariant 1728 if it is an elliptic curve. -/ -def ofJ1728 : WeierstrassCurve R := - ⟨0, 0, 0, 1, 0⟩ - -lemma ofJ1728_c₄ : (ofJ1728 R).c₄ = -48 := by - rw [ofJ1728, c₄, b₂, b₄] - norm_num1 - -lemma ofJ1728_Δ : (ofJ1728 R).Δ = -64 := by - rw [ofJ1728, Δ, b₂, b₄, b₆, b₈] - norm_num1 - -variable {R} (j : R) - -/-- The Weierstrass curve $Y^2 + (j - 1728)XY = X^3 - 36(j - 1728)^3X - (j - 1728)^5$. -It is of j-invariant j if it is an elliptic curve. -/ -def ofJ : WeierstrassCurve R := - ⟨j - 1728, 0, 0, -36 * (j - 1728) ^ 3, -(j - 1728) ^ 5⟩ - -lemma ofJ_c₄ : (ofJ j).c₄ = j * (j - 1728) ^ 3 := by - simp only [ofJ, c₄, b₂, b₄] - ring1 - -lemma ofJ_Δ : (ofJ j).Δ = j ^ 2 * (j - 1728) ^ 9 := by - simp only [ofJ, Δ, b₂, b₄, b₆, b₈] - ring1 - -end ModelsWithJ - -end WeierstrassCurve - /-! ## Elliptic curves -/ -/-- An elliptic curve over a commutative ring. Note that this definition is only mathematically +-- TODO: change to `protected abbrev IsElliptic := IsUnit W.Δ` once #17458 is merged +/-- `WeierstrassCurve.IsElliptic` is a typeclass which asserts that a Weierstrass curve is an +elliptic curve: that its discriminant is a unit. Note that this definition is only mathematically accurate for certain rings whose Picard group has trivial 12-torsion, such as a field or a PID. -/ -structure EllipticCurve (R : Type u) [CommRing R] extends WeierstrassCurve R where - /-- The discriminant `Δ'` of an elliptic curve over `R`, which is given as a unit in `R`. -/ - Δ' : Rˣ - /-- The discriminant of `E` is equal to the discriminant of `E` as a Weierstrass curve. -/ - coe_Δ' : Δ' = toWeierstrassCurve.Δ +@[mk_iff] +protected class IsElliptic : Prop where + isUnit : IsUnit W.Δ -namespace EllipticCurve +variable [W.IsElliptic] -variable {R : Type u} [CommRing R] +lemma isUnit_Δ : IsUnit W.Δ := IsElliptic.isUnit -theorem toWeierstrassCurve_injective : Function.Injective (toWeierstrassCurve (R := R)) - | ⟨x1, _, x3⟩, ⟨y1, _, y3⟩, h => by - change x1 = y1 at h - congr - exact Units.ext (by rw [x3, y3, h]) +/-- The discriminant `Δ'` of an elliptic curve over `R`, which is given as a unit in `R`. +Note that to prove two equal elliptic curves have the same `Δ'`, you need to use `simp_rw`, +as `rw` cannot transfer instance `WeierstrassCurve.IsElliptic` automatically. -/ +noncomputable def Δ' : Rˣ := W.isUnit_Δ.unit -@[ext] -theorem ext {x y : EllipticCurve R} (h₁ : x.a₁ = y.a₁) (h₂ : x.a₂ = y.a₂) (h₃ : x.a₃ = y.a₃) - (h₄ : x.a₄ = y.a₄) (h₆ : x.a₆ = y.a₆) : x = y := - toWeierstrassCurve_injective (WeierstrassCurve.ext h₁ h₂ h₃ h₄ h₆) - -variable (E : EllipticCurve R) +/-- The discriminant `Δ'` of an elliptic curve is equal to the +discriminant `Δ` of it as a Weierstrass curve. -/ +@[simp] +lemma coe_Δ' : W.Δ' = W.Δ := rfl -/-- The j-invariant `j` of an elliptic curve, which is invariant under isomorphisms over `R`. -/ -def j : R := - E.Δ'⁻¹ * E.c₄ ^ 3 +/-- The j-invariant `j` of an elliptic curve, which is invariant under isomorphisms over `R`. +Note that to prove two equal elliptic curves have the same `j`, you need to use `simp_rw`, +as `rw` cannot transfer instance `WeierstrassCurve.IsElliptic` automatically. -/ +noncomputable def j : R := + W.Δ'⁻¹ * W.c₄ ^ 3 -/-- A variant of `EllipticCurve.j_eq_zero_iff` without assuming a reduced ring. -/ -lemma j_eq_zero_iff' : E.j = 0 ↔ E.c₄ ^ 3 = 0 := by +/-- A variant of `WeierstrassCurve.j_eq_zero_iff` without assuming a reduced ring. -/ +lemma j_eq_zero_iff' : W.j = 0 ↔ W.c₄ ^ 3 = 0 := by rw [j, Units.mul_right_eq_zero] -lemma j_eq_zero (h : E.c₄ = 0) : E.j = 0 := by +lemma j_eq_zero (h : W.c₄ = 0) : W.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 +lemma j_eq_zero_iff [IsReduced R] : W.j = 0 ↔ W.c₄ = 0 := by rw [j_eq_zero_iff', IsReduced.pow_eq_zero_iff three_ne_zero] section CharTwo variable [CharP R 2] -lemma j_of_char_two : E.j = E.Δ'⁻¹ * E.a₁ ^ 12 := by - rw [j, E.c₄_of_char_two, ← pow_mul] +lemma j_of_char_two : W.j = W.Δ'⁻¹ * W.a₁ ^ 12 := by + rw [j, W.c₄_of_char_two, ← pow_mul] -/-- A variant of `EllipticCurve.j_eq_zero_iff_of_char_two` without assuming a reduced ring. -/ -lemma j_eq_zero_iff_of_char_two' : E.j = 0 ↔ E.a₁ ^ 12 = 0 := by +/-- A variant of `WeierstrassCurve.j_eq_zero_iff_of_char_two` without assuming a reduced ring. -/ +lemma j_eq_zero_iff_of_char_two' : W.j = 0 ↔ W.a₁ ^ 12 = 0 := by rw [j_of_char_two, Units.mul_right_eq_zero] -lemma j_eq_zero_of_char_two (h : E.a₁ = 0) : E.j = 0 := by +lemma j_eq_zero_of_char_two (h : W.a₁ = 0) : W.j = 0 := by rw [j_eq_zero_iff_of_char_two', h, zero_pow (Nat.succ_ne_zero _)] -lemma j_eq_zero_iff_of_char_two [IsReduced R] : E.j = 0 ↔ E.a₁ = 0 := by +lemma j_eq_zero_iff_of_char_two [IsReduced R] : W.j = 0 ↔ W.a₁ = 0 := by rw [j_eq_zero_iff_of_char_two', IsReduced.pow_eq_zero_iff (Nat.succ_ne_zero _)] end CharTwo @@ -469,24 +414,26 @@ section CharThree variable [CharP R 3] -lemma j_of_char_three : E.j = E.Δ'⁻¹ * E.b₂ ^ 6 := by - rw [j, E.c₄_of_char_three, ← pow_mul] +lemma j_of_char_three : W.j = W.Δ'⁻¹ * W.b₂ ^ 6 := by + rw [j, W.c₄_of_char_three, ← pow_mul] -/-- A variant of `EllipticCurve.j_eq_zero_iff_of_char_three` without assuming a reduced ring. -/ -lemma j_eq_zero_iff_of_char_three' : E.j = 0 ↔ E.b₂ ^ 6 = 0 := by +/-- A variant of `WeierstrassCurve.j_eq_zero_iff_of_char_three` without assuming a reduced ring. -/ +lemma j_eq_zero_iff_of_char_three' : W.j = 0 ↔ W.b₂ ^ 6 = 0 := by rw [j_of_char_three, Units.mul_right_eq_zero] -lemma j_eq_zero_of_char_three (h : E.b₂ = 0) : E.j = 0 := by +lemma j_eq_zero_of_char_three (h : W.b₂ = 0) : W.j = 0 := by rw [j_eq_zero_iff_of_char_three', h, zero_pow (Nat.succ_ne_zero _)] -lemma j_eq_zero_iff_of_char_three [IsReduced R] : E.j = 0 ↔ E.b₂ = 0 := by +lemma j_eq_zero_iff_of_char_three [IsReduced R] : W.j = 0 ↔ W.b₂ = 0 := by rw [j_eq_zero_iff_of_char_three', IsReduced.pow_eq_zero_iff (Nat.succ_ne_zero _)] end CharThree -lemma twoTorsionPolynomial_disc_ne_zero [Nontrivial R] [Invertible (2 : R)] : - E.twoTorsionPolynomial.disc ≠ 0 := - E.toWeierstrassCurve.twoTorsionPolynomial_disc_ne_zero <| E.coe_Δ' ▸ E.Δ'.isUnit +-- TODO: this is defeq to `twoTorsionPolynomial_disc_ne_zero` once #17458 is merged, +-- TODO: consider removing/rephrasing this result +lemma twoTorsionPolynomial_disc_ne_zero_of_isElliptic [Nontrivial R] (hu : IsUnit (2 : R)) : + W.twoTorsionPolynomial.disc ≠ 0 := + W.twoTorsionPolynomial_disc_ne_zero hu W.isUnit_Δ section BaseChange @@ -494,155 +441,31 @@ section BaseChange variable {A : Type v} [CommRing A] (φ : R →+* A) --- Porting note: was just `@[simps]` -/-- The elliptic curve mapped over a ring homomorphism `φ : R →+* A`. -/ -@[simps (config := { rhsMd := .default }) a₁ a₂ a₃ a₄ a₆ Δ' toWeierstrassCurve] -def map : EllipticCurve A := - ⟨E.toWeierstrassCurve.map φ, Units.map φ E.Δ', by simp only [Units.coe_map, coe_Δ', E.map_Δ]; rfl⟩ - -variable (A) +instance : (W.map φ).IsElliptic := by + simp only [isElliptic_iff, map_Δ, W.isUnit_Δ.map] -/-- The elliptic curve base changed to an algebra `A` over `R`. -/ -abbrev baseChange [Algebra R A] : EllipticCurve A := - E.map <| algebraMap R A +set_option linter.docPrime false in +lemma coe_map_Δ' : (W.map φ).Δ' = φ W.Δ' := by + rw [coe_Δ', map_Δ, coe_Δ'] -variable {A} +set_option linter.docPrime false in +@[simp] +lemma map_Δ' : (W.map φ).Δ' = Units.map φ W.Δ' := by + ext + exact W.coe_map_Δ' φ -lemma coe_map_Δ' : (E.map φ).Δ' = φ E.Δ' := - rfl +set_option linter.docPrime false in +lemma coe_inv_map_Δ' : (W.map φ).Δ'⁻¹ = φ ↑W.Δ'⁻¹ := by + simp -lemma coe_inv_map_Δ' : (E.map φ).Δ'⁻¹ = φ ↑E.Δ'⁻¹ := - rfl +set_option linter.docPrime false in +lemma inv_map_Δ' : (W.map φ).Δ'⁻¹ = Units.map φ W.Δ'⁻¹ := by + simp @[simp] -lemma map_j : (E.map φ).j = φ E.j := by - simp [j, map, E.map_c₄] - -lemma map_injective {φ : R →+* A} (hφ : Function.Injective φ) : - Function.Injective <| map (φ := φ) := fun _ _ h => by - rcases mk.inj h with ⟨h1, _⟩ - rcases WeierstrassCurve.mk.inj h1 with ⟨_, _, _, _, _⟩ - ext <;> apply_fun _ using hφ <;> assumption +lemma map_j : (W.map φ).j = φ W.j := by + rw [j, coe_inv_map_Δ', map_c₄, j, map_mul, map_pow] end BaseChange -section ModelsWithJ - -/-! ### Models with prescribed j-invariant -/ - -variable (R) - -/-- When 3 is invertible, $Y^2 + Y = X^3$ is an elliptic curve. -It is of j-invariant 0 (see `EllipticCurve.ofJ0_j`). -/ -def ofJ0 [Invertible (3 : R)] : EllipticCurve R := - have := invertibleNeg (3 ^ 3 : R) - ⟨WeierstrassCurve.ofJ0 R, unitOfInvertible (-3 ^ 3 : R), - by rw [val_unitOfInvertible, WeierstrassCurve.ofJ0_Δ R]; norm_num1⟩ - -lemma ofJ0_j [Invertible (3 : R)] : (ofJ0 R).j = 0 := by - simp only [j, ofJ0, WeierstrassCurve.ofJ0_c₄] - ring1 - -/-- When 2 is invertible, $Y^2 = X^3 + X$ is an elliptic curve. -It is of j-invariant 1728 (see `EllipticCurve.ofJ1728_j`). -/ -def ofJ1728 [Invertible (2 : R)] : EllipticCurve R := - have := invertibleNeg (2 ^ 6 : R) - ⟨WeierstrassCurve.ofJ1728 R, unitOfInvertible (-2 ^ 6 : R), - by rw [val_unitOfInvertible, WeierstrassCurve.ofJ1728_Δ R]; norm_num1⟩ - -lemma ofJ1728_j [Invertible (2 : R)] : (ofJ1728 R).j = 1728 := by - field_simp [j, ofJ1728, @val_unitOfInvertible _ _ _ <| invertibleNeg _, - WeierstrassCurve.ofJ1728_c₄] - norm_num1 - -variable {R} - -/-- When j and j - 1728 are both invertible, -$Y^2 + (j - 1728)XY = X^3 - 36(j - 1728)^3X - (j - 1728)^5$ is an elliptic curve. -It is of j-invariant j (see `EllipticCurve.ofJ'_j`). -/ -def ofJ' (j : R) [Invertible j] [Invertible (j - 1728)] : EllipticCurve R := - have := invertibleMul (j ^ 2) ((j - 1728) ^ 9) - ⟨WeierstrassCurve.ofJ j, unitOfInvertible <| j ^ 2 * (j - 1728) ^ 9, - (WeierstrassCurve.ofJ_Δ j).symm⟩ - -lemma ofJ'_j (j : R) [Invertible j] [Invertible (j - 1728)] : (ofJ' j).j = j := by - field_simp [EllipticCurve.j, ofJ', @val_unitOfInvertible _ _ _ <| invertibleMul .., - WeierstrassCurve.ofJ_c₄] - ring1 - -variable {F : Type u} [Field F] (j : F) - -private lemma two_or_three_ne_zero : (2 : F) ≠ 0 ∨ (3 : F) ≠ 0 := - ne_zero_or_ne_zero_of_nat_coprime <| by decide - -variable [DecidableEq F] - -/-- For any element j of a field $F$, there exists an elliptic curve over $F$ -with j-invariant equal to j (see `EllipticCurve.ofJ_j`). -Its coefficients are given explicitly (see `EllipticCurve.ofJ0`, `EllipticCurve.ofJ1728` -and `EllipticCurve.ofJ'`). -/ -def ofJ : EllipticCurve F := - if h0 : j = 0 then - if h3 : (3 : F) = 0 then @ofJ1728 _ _ <| invertibleOfNonzero <| - two_or_three_ne_zero.neg_resolve_right h3 - else @ofJ0 _ _ <| invertibleOfNonzero h3 - else if h1728 : j = 1728 then - @ofJ1728 _ _ <| invertibleOfNonzero fun h => h0 <| - by rw [h1728, show (1728 : F) = 2 * 864 by norm_num1, h, zero_mul] - else @ofJ' _ _ j (invertibleOfNonzero h0) (invertibleOfNonzero <| sub_ne_zero_of_ne h1728) - -lemma ofJ_0_of_three_ne_zero [h3 : NeZero (3 : F)] : - ofJ 0 = @ofJ0 _ _ (invertibleOfNonzero h3.out) := by - rw [ofJ, dif_pos rfl, dif_neg h3.out] - -lemma ofJ_0_of_three_eq_zero (h3 : (3 : F) = 0) : - ofJ 0 = @ofJ1728 _ _ (invertibleOfNonzero <| two_or_three_ne_zero.neg_resolve_right h3) := by - rw [ofJ, dif_pos rfl, dif_pos h3] - -lemma ofJ_0_of_two_eq_zero (h2 : (2 : F) = 0) : - ofJ 0 = @ofJ0 _ _ (invertibleOfNonzero <| two_or_three_ne_zero.neg_resolve_left h2) := - have := neZero_iff.2 <| two_or_three_ne_zero.neg_resolve_left h2 - ofJ_0_of_three_ne_zero - -lemma ofJ_1728_of_three_eq_zero (h3 : (3 : F) = 0) : - ofJ 1728 = @ofJ1728 _ _ (invertibleOfNonzero <| two_or_three_ne_zero.neg_resolve_right h3) := by - rw [ofJ, dif_pos <| by rw [show (1728 : F) = 3 * 576 by norm_num1, h3, zero_mul], dif_pos h3] - -lemma ofJ_1728_of_two_ne_zero [h2 : NeZero (2 : F)] : - ofJ 1728 = @ofJ1728 _ _ (invertibleOfNonzero h2.out) := by - by_cases h3 : (3 : F) = 0 - · exact ofJ_1728_of_three_eq_zero h3 - · have h : (1728 : F) ≠ 0 := fun h => or_iff_not_and_not.mp - (mul_eq_zero.mp <| by rwa [show 2 ^ 6 * 3 ^ 3 = (1728 : F) by norm_num1]) - ⟨pow_ne_zero 6 h2.out, pow_ne_zero 3 h3⟩ - rw [ofJ, dif_neg h, dif_pos rfl] - -lemma ofJ_1728_of_two_eq_zero (h2 : (2 : F) = 0) : - ofJ 1728 = @ofJ0 _ _ (invertibleOfNonzero <| two_or_three_ne_zero.neg_resolve_left h2) := by - rw [ofJ, dif_pos <| by rw [show (1728 : F) = 2 * 864 by norm_num1, h2, zero_mul], dif_neg] - -lemma ofJ_ne_0_ne_1728 (h0 : j ≠ 0) (h1728 : j ≠ 1728) : ofJ j = - @ofJ' _ _ j (invertibleOfNonzero h0) (invertibleOfNonzero <| sub_ne_zero_of_ne h1728) := by - rw [ofJ, dif_neg h0, dif_neg h1728] - -lemma ofJ_j : (ofJ j).j = j := by - by_cases h0 : j = 0 - · by_cases h3 : (3 : F) = 0 - · rw [h0, ofJ_0_of_three_eq_zero h3, - @ofJ1728_j _ _ <| invertibleOfNonzero <| two_or_three_ne_zero.neg_resolve_right h3, - show (1728 : F) = 3 * 576 by norm_num1, h3, zero_mul] - · rw [h0, ofJ_0_of_three_ne_zero (h3 := neZero_iff.2 h3), @ofJ0_j _ _ <| invertibleOfNonzero h3] - · by_cases h1728 : j = 1728 - · have h2 : (2 : F) ≠ 0 := - fun h => h0 <| by rw [h1728, show (1728 : F) = 2 * 864 by norm_num1, h, zero_mul] - rw [h1728, ofJ_1728_of_two_ne_zero (h2 := neZero_iff.2 h2), - @ofJ1728_j _ _ <| invertibleOfNonzero h2] - · rw [ofJ_ne_0_ne_1728 j h0 h1728, - @ofJ'_j _ _ _ (invertibleOfNonzero h0) (invertibleOfNonzero <| sub_ne_zero_of_ne h1728)] - -instance instInhabitedEllipticCurve : Inhabited <| EllipticCurve F := - ⟨ofJ 37⟩ - -end ModelsWithJ - -end EllipticCurve +end WeierstrassCurve From c6be01e3534aa6acbba1701ffa9131a59be5ce51 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Fri, 29 Nov 2024 23:24:21 +0000 Subject: [PATCH 422/829] chore: split out map_dvd (#19565) --- Mathlib.lean | 1 + Mathlib/Algebra/Divisibility/Basic.lean | 17 ---------- Mathlib/Algebra/Divisibility/Hom.lean | 34 ++++++++++++++++++++ Mathlib/Algebra/Prime/Lemmas.lean | 1 + Mathlib/Algebra/Ring/Divisibility/Basic.lean | 2 +- Mathlib/Algebra/Ring/Hom/Basic.lean | 2 +- Mathlib/Data/Nat/Cast/Basic.lean | 2 +- 7 files changed, 39 insertions(+), 20 deletions(-) create mode 100644 Mathlib/Algebra/Divisibility/Hom.lean diff --git a/Mathlib.lean b/Mathlib.lean index 027c728331798..8f03dcdf54086 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -202,6 +202,7 @@ import Mathlib.Algebra.DirectSum.LinearMap import Mathlib.Algebra.DirectSum.Module import Mathlib.Algebra.DirectSum.Ring import Mathlib.Algebra.Divisibility.Basic +import Mathlib.Algebra.Divisibility.Hom import Mathlib.Algebra.Divisibility.Prod import Mathlib.Algebra.Divisibility.Units import Mathlib.Algebra.DualNumber diff --git a/Mathlib/Algebra/Divisibility/Basic.lean b/Mathlib/Algebra/Divisibility/Basic.lean index 7e62a5bd031dc..c51368437a942 100644 --- a/Mathlib/Algebra/Divisibility/Basic.lean +++ b/Mathlib/Algebra/Divisibility/Basic.lean @@ -5,7 +5,6 @@ Authors: Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Amelia Livingston, Neil Strickland, Aaron Anderson -/ import Mathlib.Algebra.Group.Basic -import Mathlib.Algebra.Group.Hom.Defs import Mathlib.Tactic.Common /-! @@ -83,22 +82,6 @@ alias Dvd.dvd.mul_right := dvd_mul_of_dvd_left theorem dvd_of_mul_right_dvd (h : a * b ∣ c) : a ∣ c := (dvd_mul_right a b).trans h -section map_dvd - -variable {M N : Type*} - -theorem map_dvd [Semigroup M] [Semigroup N] {F : Type*} [FunLike F M N] [MulHomClass F M N] - (f : F) {a b} : a ∣ b → f a ∣ f b - | ⟨c, h⟩ => ⟨f c, h.symm ▸ map_mul f a c⟩ - -theorem MulHom.map_dvd [Semigroup M] [Semigroup N] (f : M →ₙ* N) {a b} : a ∣ b → f a ∣ f b := - _root_.map_dvd f - -theorem MonoidHom.map_dvd [Monoid M] [Monoid N] (f : M →* N) {a b} : a ∣ b → f a ∣ f b := - _root_.map_dvd f - -end map_dvd - /-- An element `a` in a semigroup is primal if whenever `a` is a divisor of `b * c`, it can be factored as the product of a divisor of `b` and a divisor of `c`. -/ def IsPrimal (a : α) : Prop := ∀ ⦃b c⦄, a ∣ b * c → ∃ a₁ a₂, a₁ ∣ b ∧ a₂ ∣ c ∧ a = a₁ * a₂ diff --git a/Mathlib/Algebra/Divisibility/Hom.lean b/Mathlib/Algebra/Divisibility/Hom.lean new file mode 100644 index 0000000000000..36e3874d14974 --- /dev/null +++ b/Mathlib/Algebra/Divisibility/Hom.lean @@ -0,0 +1,34 @@ +/- +Copyright (c) 2014 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Amelia Livingston, Yury Kudryashov, +Neil Strickland, Aaron Anderson +-/ +import Mathlib.Algebra.Divisibility.Basic +import Mathlib.Algebra.Group.Hom.Defs + +/-! +# Mapping divisibility across multiplication-preserving homomorphisms + +## Main definitions + + * `map_dvd` + +## Tags + +divisibility, divides +-/ + +attribute [local simp] mul_assoc mul_comm mul_left_comm + +variable {M N : Type*} + +theorem map_dvd [Semigroup M] [Semigroup N] {F : Type*} [FunLike F M N] [MulHomClass F M N] + (f : F) {a b} : a ∣ b → f a ∣ f b + | ⟨c, h⟩ => ⟨f c, h.symm ▸ map_mul f a c⟩ + +theorem MulHom.map_dvd [Semigroup M] [Semigroup N] (f : M →ₙ* N) {a b} : a ∣ b → f a ∣ f b := + _root_.map_dvd f + +theorem MonoidHom.map_dvd [Monoid M] [Monoid N] (f : M →* N) {a b} : a ∣ b → f a ∣ f b := + _root_.map_dvd f diff --git a/Mathlib/Algebra/Prime/Lemmas.lean b/Mathlib/Algebra/Prime/Lemmas.lean index 0211c3bbc8c5e..1df2748e62898 100644 --- a/Mathlib/Algebra/Prime/Lemmas.lean +++ b/Mathlib/Algebra/Prime/Lemmas.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Jens Wagemaker -/ +import Mathlib.Algebra.Divisibility.Hom import Mathlib.Algebra.Group.Commute.Units import Mathlib.Algebra.Group.Even import Mathlib.Algebra.Group.Units.Equiv diff --git a/Mathlib/Algebra/Ring/Divisibility/Basic.lean b/Mathlib/Algebra/Ring/Divisibility/Basic.lean index dbe11489b614e..1f2823ca80b9c 100644 --- a/Mathlib/Algebra/Ring/Divisibility/Basic.lean +++ b/Mathlib/Algebra/Ring/Divisibility/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2014 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Yury Kudryashov, Neil Strickland -/ -import Mathlib.Algebra.Divisibility.Basic +import Mathlib.Algebra.Divisibility.Hom import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.Algebra.Ring.Defs diff --git a/Mathlib/Algebra/Ring/Hom/Basic.lean b/Mathlib/Algebra/Ring/Hom/Basic.lean index 66e709402768d..90d26a14ce0cb 100644 --- a/Mathlib/Algebra/Ring/Hom/Basic.lean +++ b/Mathlib/Algebra/Ring/Hom/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Amelia Livingston. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Amelia Livingston, Jireh Loreaux -/ -import Mathlib.Algebra.Divisibility.Basic +import Mathlib.Algebra.Divisibility.Hom import Mathlib.Algebra.GroupWithZero.InjSurj import Mathlib.Algebra.Ring.Hom.Defs import Mathlib.Data.Set.Basic diff --git a/Mathlib/Data/Nat/Cast/Basic.lean b/Mathlib/Data/Nat/Cast/Basic.lean index c494b68c05067..1b66e7d26222f 100644 --- a/Mathlib/Data/Nat/Cast/Basic.lean +++ b/Mathlib/Data/Nat/Cast/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2014 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Mathlib.Algebra.Divisibility.Basic +import Mathlib.Algebra.Divisibility.Hom import Mathlib.Algebra.Group.Even import Mathlib.Algebra.Group.TypeTags.Hom import Mathlib.Algebra.Ring.Hom.Defs From 9bdc1af76cf75d960f42c29f1c09917785066c8b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Sat, 30 Nov 2024 00:02:59 +0000 Subject: [PATCH 423/829] feat(Algebra/Homology/Embedding): the left homology of an extension of homological complexes (#18502) Given an embedding `e : c.Embedding c'`, `K : HomologicalComplex C c`, and degrees such that `e.f j = j'`, we construct a left homology data for `K.extend e` in degree `j'` from a left homology data for `K` in degree `j`. --- Mathlib.lean | 1 + .../Homology/Embedding/ExtendHomology.lean | 143 ++++++++++++++++++ .../CategoryTheory/Limits/Shapes/Kernels.lean | 46 ++++++ 3 files changed, 190 insertions(+) create mode 100644 Mathlib/Algebra/Homology/Embedding/ExtendHomology.lean diff --git a/Mathlib.lean b/Mathlib.lean index 8f03dcdf54086..85bd8c077ab33 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -406,6 +406,7 @@ import Mathlib.Algebra.Homology.DifferentialObject import Mathlib.Algebra.Homology.Embedding.Basic import Mathlib.Algebra.Homology.Embedding.Boundary import Mathlib.Algebra.Homology.Embedding.Extend +import Mathlib.Algebra.Homology.Embedding.ExtendHomology import Mathlib.Algebra.Homology.Embedding.HomEquiv import Mathlib.Algebra.Homology.Embedding.IsSupported import Mathlib.Algebra.Homology.Embedding.Restriction diff --git a/Mathlib/Algebra/Homology/Embedding/ExtendHomology.lean b/Mathlib/Algebra/Homology/Embedding/ExtendHomology.lean new file mode 100644 index 0000000000000..123c91edcc3e7 --- /dev/null +++ b/Mathlib/Algebra/Homology/Embedding/ExtendHomology.lean @@ -0,0 +1,143 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.Algebra.Homology.Embedding.Extend +import Mathlib.Algebra.Homology.ShortComplex.HomologicalComplex + +/-! +# Homology of the extension of an homological complex + +Given an embedding `e : c.Embedding c'` and `K : HomologicalComplex C c`, we shall +compute the homology of `K.extend e`. In degrees that are not in the image of `e.f`, +the homology is obviously zero. When `e.f j = j`, we shall construct an isomorphism +`(K.extend e).homology j' ≅ K.homology j` (TODO). + +-/ + +open CategoryTheory Limits Category + +namespace HomologicalComplex + +variable {ι ι' : Type*} {c : ComplexShape ι} {c' : ComplexShape ι'} + {C : Type*} [Category C] [HasZeroMorphisms C] [HasZeroObject C] + +variable (K L M : HomologicalComplex C c) (φ : K ⟶ L) (φ' : L ⟶ M) (e : c.Embedding c') + +namespace extend + +section HomologyData + +variable {i j k : ι} {i' j' k' : ι'} (hj' : e.f j = j') + (hi : c.prev j = i) (hi' : c'.prev j' = i') (hk : c.next j = k) (hk' : c'.next j' = k') + +include hk hk' in +lemma comp_d_eq_zero_iff ⦃W : C⦄ (φ : W ⟶ K.X j) : + φ ≫ K.d j k = 0 ↔ φ ≫ (K.extendXIso e hj').inv ≫ (K.extend e).d j' k' = 0 := by + by_cases hjk : c.Rel j k + · have hk' : e.f k = k' := by rw [← hk', ← hj', c'.next_eq' (e.rel hjk)] + rw [K.extend_d_eq e hj' hk', Iso.inv_hom_id_assoc, + ← cancel_mono (K.extendXIso e hk').inv, zero_comp, assoc] + · simp only [K.shape _ _ hjk, comp_zero, true_iff] + rw [K.extend_d_from_eq_zero e j' k' j hj', comp_zero, comp_zero] + rw [hk] + exact hjk + +namespace leftHomologyData + +variable (cone : KernelFork (K.d j k)) (hcone : IsLimit cone) + +/-- The kernel fork of `(K.extend e).d j' k'` that is deduced from a kernel +fork of `K.d j k `. -/ +@[simp] +noncomputable def kernelFork : KernelFork ((K.extend e).d j' k') := + KernelFork.ofι (cone.ι ≫ (extendXIso K e hj').inv) + (by rw [assoc, ← comp_d_eq_zero_iff K e hj' hk hk' cone.ι, cone.condition]) + +/-- The limit kernel fork of `(K.extend e).d j' k'` that is deduced from a limit +kernel fork of `K.d j k `. -/ +noncomputable def isLimitKernelFork : IsLimit (kernelFork K e hj' hk hk' cone) := + KernelFork.isLimitOfIsLimitOfIff hcone ((K.extend e).d j' k') + (extendXIso K e hj').symm (comp_d_eq_zero_iff K e hj' hk hk') + +variable (cocone : CokernelCofork (hcone.lift (KernelFork.ofι (K.d i j) (K.d_comp_d i j k)))) + (hcocone : IsColimit cocone) + +include hi hi' hcone in +/-- Auxiliary lemma for `lift_d_comp_eq_zero_iff`. -/ +lemma lift_d_comp_eq_zero_iff' ⦃W : C⦄ (f' : K.X i ⟶ cone.pt) + (hf' : f' ≫ cone.ι = K.d i j) + (f'' : (K.extend e).X i' ⟶ cone.pt) + (hf'' : f'' ≫ cone.ι ≫ (extendXIso K e hj').inv = (K.extend e).d i' j') + (φ : cone.pt ⟶ W) : + f' ≫ φ = 0 ↔ f'' ≫ φ = 0 := by + by_cases hij : c.Rel i j + · have hi'' : e.f i = i' := by rw [← hi', ← hj', c'.prev_eq' (e.rel hij)] + have : (K.extendXIso e hi'').hom ≫ f' = f'' := by + apply Fork.IsLimit.hom_ext hcone + rw [assoc, hf', ← cancel_mono (extendXIso K e hj').inv, assoc, assoc, hf'', + K.extend_d_eq e hi'' hj'] + rw [← cancel_epi (K.extendXIso e hi'').hom, comp_zero, ← this, assoc] + · have h₁ : f' = 0 := by + apply Fork.IsLimit.hom_ext hcone + simp only [zero_comp, hf', K.shape _ _ hij] + have h₂ : f'' = 0 := by + apply Fork.IsLimit.hom_ext hcone + dsimp + rw [← cancel_mono (extendXIso K e hj').inv, assoc, hf'', zero_comp, zero_comp, + K.extend_d_to_eq_zero e i' j' j hj'] + rw [hi] + exact hij + simp [h₁, h₂] + +include hi hi' in +lemma lift_d_comp_eq_zero_iff ⦃W : C⦄ (φ : cone.pt ⟶ W) : + hcone.lift (KernelFork.ofι (K.d i j) (K.d_comp_d i j k)) ≫ φ = 0 ↔ + ((isLimitKernelFork K e hj' hk hk' cone hcone).lift + (KernelFork.ofι ((K.extend e).d i' j') (d_comp_d _ _ _ _))) ≫ φ = 0 := + lift_d_comp_eq_zero_iff' K e hj' hi hi' cone hcone _ (hcone.fac _ _) _ + (IsLimit.fac _ _ WalkingParallelPair.zero) _ + +/-- Auxiliary definition for `extend.leftHomologyData`. -/ +noncomputable def cokernelCofork : + CokernelCofork ((isLimitKernelFork K e hj' hk hk' cone hcone).lift + (KernelFork.ofι ((K.extend e).d i' j') (d_comp_d _ _ _ _))) := + CokernelCofork.ofπ cocone.π (by + rw [← lift_d_comp_eq_zero_iff K e hj' hi hi' hk hk' cone hcone] + exact cocone.condition) + +/-- Auxiliary definition for `extend.leftHomologyData`. -/ +noncomputable def isColimitCokernelCofork : + IsColimit (cokernelCofork K e hj' hi hi' hk hk' cone hcone cocone) := + CokernelCofork.isColimitOfIsColimitOfIff' hcocone _ + (lift_d_comp_eq_zero_iff K e hj' hi hi' hk hk' cone hcone) + +end leftHomologyData + +open leftHomologyData in +/-- The left homology data of `(K.extend e).sc' i' j' k'` that is deduced +from a left homology data of `K.sc' i j k`. -/ +@[simps] +noncomputable def leftHomologyData (h : (K.sc' i j k).LeftHomologyData) : + ((K.extend e).sc' i' j' k').LeftHomologyData where + K := h.K + H := h.H + i := h.i ≫ (extendXIso K e hj').inv + π := h.π + wi := by + dsimp + rw [assoc, ← comp_d_eq_zero_iff K e hj' hk hk'] + exact h.wi + hi := isLimitKernelFork K e hj' hk hk' _ h.hi + wπ := by + dsimp + rw [← lift_d_comp_eq_zero_iff K e hj' hi hi' hk hk' _ h.hi] + exact h.wπ + hπ := isColimitCokernelCofork K e hj' hi hi' hk hk' _ h.hi _ h.hπ + +end HomologyData + +end extend + +end HomologicalComplex diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean b/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean index f03cd551569a5..202f266d2d66f 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Kernels.lean @@ -200,6 +200,29 @@ lemma KernelFork.IsLimit.isIso_ι {X Y : C} {f : X ⟶ Y} (c : KernelFork f) infer_instance exact IsIso.of_isIso_comp_left e.inv c.ι +/-- If `c` is a limit kernel fork for `g : X ⟶ Y`, `e : X ≅ X'` and `g' : X' ⟶ Y` is a morphism, +then there is a limit kernel fork for `g'` with the same point as `c` if for any +morphism `φ : W ⟶ X`, there is an equivalence `φ ≫ g = 0 ↔ φ ≫ e.hom ≫ g' = 0`. -/ +def KernelFork.isLimitOfIsLimitOfIff {X Y : C} {g : X ⟶ Y} {c : KernelFork g} (hc : IsLimit c) + {X' Y' : C} (g' : X' ⟶ Y') (e : X ≅ X') + (iff : ∀ ⦃W : C⦄ (φ : W ⟶ X), φ ≫ g = 0 ↔ φ ≫ e.hom ≫ g' = 0) : + IsLimit (KernelFork.ofι (f := g') (c.ι ≫ e.hom) (by simp [← iff])) := + KernelFork.IsLimit.ofι _ _ + (fun s hs ↦ hc.lift (KernelFork.ofι (ι := s ≫ e.inv) + (by rw [iff, Category.assoc, Iso.inv_hom_id_assoc, hs]))) + (fun s hs ↦ by simp [← cancel_mono e.inv]) + (fun s hs m hm ↦ Fork.IsLimit.hom_ext hc (by simpa [← cancel_mono e.hom] using hm)) + +/-- If `c` is a limit kernel fork for `g : X ⟶ Y`, and `g' : X ⟶ Y'` is a another morphism, +then there is a limit kernel fork for `g'` with the same point as `c` if for any +morphism `φ : W ⟶ X`, there is an equivalence `φ ≫ g = 0 ↔ φ ≫ g' = 0`. -/ +def KernelFork.isLimitOfIsLimitOfIff' {X Y : C} {g : X ⟶ Y} {c : KernelFork g} (hc : IsLimit c) + {Y' : C} (g' : X ⟶ Y') + (iff : ∀ ⦃W : C⦄ (φ : W ⟶ X), φ ≫ g = 0 ↔ φ ≫ g' = 0) : + IsLimit (KernelFork.ofι (f := g') c.ι (by simp [← iff])) := + IsLimit.ofIsoLimit (isLimitOfIsLimitOfIff hc g' (Iso.refl _) (by simpa using iff)) + (Fork.ext (Iso.refl _)) + end namespace KernelFork @@ -632,6 +655,29 @@ lemma CokernelCofork.IsColimit.isIso_π {X Y : C} {f : X ⟶ Y} (c : CokernelCof infer_instance exact IsIso.of_isIso_comp_right c.π e.hom +/-- If `c` is a colimit cokernel cofork for `f : X ⟶ Y`, `e : Y ≅ Y'` and `f' : X' ⟶ Y` is a +morphism, then there is a colimit cokernel cofork for `f'` with the same point as `c` if for any +morphism `φ : Y ⟶ W`, there is an equivalence `f ≫ φ = 0 ↔ f' ≫ e.hom ≫ φ = 0`. -/ +def CokernelCofork.isColimitOfIsColimitOfIff {X Y : C} {f : X ⟶ Y} {c : CokernelCofork f} + (hc : IsColimit c) {X' Y' : C} (f' : X' ⟶ Y') (e : Y' ≅ Y) + (iff : ∀ ⦃W : C⦄ (φ : Y ⟶ W), f ≫ φ = 0 ↔ f' ≫ e.hom ≫ φ = 0) : + IsColimit (CokernelCofork.ofπ (f := f') (e.hom ≫ c.π) (by simp [← iff])) := + CokernelCofork.IsColimit.ofπ _ _ + (fun s hs ↦ hc.desc (CokernelCofork.ofπ (π := e.inv ≫ s) + (by rw [iff, e.hom_inv_id_assoc, hs]))) + (fun s hs ↦ by simp [← cancel_epi e.inv]) + (fun s hs m hm ↦ Cofork.IsColimit.hom_ext hc (by simpa [← cancel_epi e.hom] using hm)) + +/-- If `c` is a colimit cokernel cofork for `f : X ⟶ Y`, and `f' : X' ⟶ Y is another +morphism, then there is a colimit cokernel cofork for `f'` with the same point as `c` if for any +morphism `φ : Y ⟶ W`, there is an equivalence `f ≫ φ = 0 ↔ f' ≫ φ = 0`. -/ +def CokernelCofork.isColimitOfIsColimitOfIff' {X Y : C} {f : X ⟶ Y} {c : CokernelCofork f} + (hc : IsColimit c) {X' : C} (f' : X' ⟶ Y) + (iff : ∀ ⦃W : C⦄ (φ : Y ⟶ W), f ≫ φ = 0 ↔ f' ≫ φ = 0) : + IsColimit (CokernelCofork.ofπ (f := f') c.π (by simp [← iff])) := + IsColimit.ofIsoColimit (isColimitOfIsColimitOfIff hc f' (Iso.refl _) (by simpa using iff)) + (Cofork.ext (Iso.refl _)) + end namespace CokernelCofork From f33352ab9946d8ee38f17f55c04469e5befa3e6a Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Sat, 30 Nov 2024 00:03:01 +0000 Subject: [PATCH 424/829] feat(LinearAlgebra/Matrix/Permanent): smul (#19307) adding - `permanent_smul` - `permanent_updateRow_smul` - `permanent_updateColumn_smul` Co-authored-by: Moritz Firsching --- Mathlib/LinearAlgebra/Matrix/Permanent.lean | 28 +++++++++++++++++++-- 1 file changed, 26 insertions(+), 2 deletions(-) diff --git a/Mathlib/LinearAlgebra/Matrix/Permanent.lean b/Mathlib/LinearAlgebra/Matrix/Permanent.lean index fa1620053c90a..59b4a3a2b020d 100644 --- a/Mathlib/LinearAlgebra/Matrix/Permanent.lean +++ b/Mathlib/LinearAlgebra/Matrix/Permanent.lean @@ -3,9 +3,8 @@ Copyright (c) 2024 Moritz Firsching. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Moritz Firsching -/ -import Mathlib.Algebra.BigOperators.GroupWithZero.Finset import Mathlib.Data.Fintype.Perm -import Mathlib.Data.Matrix.Diagonal +import Mathlib.Data.Matrix.RowCol /-! # Permanent of a matrix @@ -83,4 +82,29 @@ theorem permanent_permute_rows (σ : Perm n) (M : Matrix n n R) : (M.submatrix id σ).permanent = M.permanent := by rw [← permanent_transpose, transpose_submatrix, permanent_permute_cols, permanent_transpose] +@[simp] +theorem permanent_smul (M : Matrix n n R) (c : R) : + permanent (c • M) = c ^ Fintype.card n * permanent M := by + simp only [permanent, smul_apply, smul_eq_mul, Finset.mul_sum] + congr + ext + rw [mul_comm] + conv in ∏ _ , c * _ => simp [mul_comm c]; + exact prod_mul_pow_card.symm + +@[simp] +theorem permanent_updateColumn_smul (M : Matrix n n R) (j : n) (c : R) (u : n → R) : + permanent (updateColumn M j <| c • u) = c * permanent (updateColumn M j u) := by + simp only [permanent, ← mul_prod_erase _ _ (mem_univ j), updateColumn_self, Pi.smul_apply, + smul_eq_mul, mul_sum, ← mul_assoc] + congr 1 with p + rw [Finset.prod_congr rfl (fun i hi ↦ ?_)] + simp only [ne_eq, ne_of_mem_erase hi, not_false_eq_true, updateColumn_ne] + +@[simp] +theorem permanent_updateRow_smul (M : Matrix n n R) (j : n) (c : R) (u : n → R) : + permanent (updateRow M j <| c • u) = c * permanent (updateRow M j u) := by + rw [← permanent_transpose, ← updateColumn_transpose, permanent_updateColumn_smul, + updateColumn_transpose, permanent_transpose] + end Matrix From 3d4e9f41345cd7e77118fad146bb07e12c8f98dd Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Sat, 30 Nov 2024 00:03:02 +0000 Subject: [PATCH 425/829] =?UTF-8?q?chore:=20change=20associativity=20of=20?= =?UTF-8?q?`+=E1=B5=A5`=20from=20`infixl`=20to=20`infixr`=20(#19321)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Zulip: https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/Associativity.20rules.20for.20SMul.20and.20VAdd --- Mathlib/Algebra/AddTorsor.lean | 18 +++++++++--------- Mathlib/Algebra/Group/Action/Defs.lean | 4 ++-- Mathlib/Algebra/Group/Operations.lean | 2 +- Mathlib/Analysis/Normed/Affine/Isometry.lean | 2 +- Mathlib/Geometry/Euclidean/Basic.lean | 7 ++++--- .../LinearAlgebra/AffineSpace/AffineEquiv.lean | 2 +- .../AffineSpace/AffineSubspace.lean | 5 +++-- 7 files changed, 21 insertions(+), 19 deletions(-) diff --git a/Mathlib/Algebra/AddTorsor.lean b/Mathlib/Algebra/AddTorsor.lean index 5c7a2b005055b..cbcdf82f7c7b7 100644 --- a/Mathlib/Algebra/AddTorsor.lean +++ b/Mathlib/Algebra/AddTorsor.lean @@ -49,7 +49,7 @@ class AddTorsor (G : outParam Type*) (P : Type*) [AddGroup G] extends AddAction /-- Torsor subtraction and addition with the same element cancels out. -/ vsub_vadd' : ∀ p₁ p₂ : P, (p₁ -ᵥ p₂ : G) +ᵥ p₂ = p₁ /-- Torsor addition and subtraction with the same element cancels out. -/ - vadd_vsub' : ∀ (g : G) (p : P), g +ᵥ p -ᵥ p = g + vadd_vsub' : ∀ (g : G) (p : P), (g +ᵥ p) -ᵥ p = g -- Porting note (https://github.com/leanprover-community/mathlib4/issues/12096): removed `nolint instance_priority`; lint not ported yet attribute [instance 100] AddTorsor.nonempty @@ -75,13 +75,13 @@ variable {G : Type*} {P : Type*} [AddGroup G] [T : AddTorsor G P] /-- Adding the result of subtracting from another point produces that point. -/ @[simp] -theorem vsub_vadd (p₁ p₂ : P) : p₁ -ᵥ p₂ +ᵥ p₂ = p₁ := +theorem vsub_vadd (p₁ p₂ : P) : (p₁ -ᵥ p₂) +ᵥ p₂ = p₁ := AddTorsor.vsub_vadd' p₁ p₂ /-- Adding a group element then subtracting the original point produces that group element. -/ @[simp] -theorem vadd_vsub (g : G) (p : P) : g +ᵥ p -ᵥ p = g := +theorem vadd_vsub (g : G) (p : P) : (g +ᵥ p) -ᵥ p = g := AddTorsor.vadd_vsub' g p /-- If the same point added to two group elements produces equal @@ -102,7 +102,7 @@ theorem vadd_right_injective (p : P) : Function.Injective ((· +ᵥ p) : G → P /-- Adding a group element to a point, then subtracting another point, produces the same result as subtracting the points then adding the group element. -/ -theorem vadd_vsub_assoc (g : G) (p₁ p₂ : P) : g +ᵥ p₁ -ᵥ p₂ = g + (p₁ -ᵥ p₂) := by +theorem vadd_vsub_assoc (g : G) (p₁ p₂ : P) : (g +ᵥ p₁) -ᵥ p₂ = g + (p₁ -ᵥ p₂) := by apply vadd_right_cancel p₂ rw [vsub_vadd, add_vadd, vsub_vadd] @@ -137,7 +137,7 @@ theorem neg_vsub_eq_vsub_rev (p₁ p₂ : P) : -(p₁ -ᵥ p₂) = p₂ -ᵥ p refine neg_eq_of_add_eq_zero_right (vadd_right_cancel p₁ ?_) rw [vsub_add_vsub_cancel, vsub_self] -theorem vadd_vsub_eq_sub_vsub (g : G) (p q : P) : g +ᵥ p -ᵥ q = g - (q -ᵥ p) := by +theorem vadd_vsub_eq_sub_vsub (g : G) (p q : P) : (g +ᵥ p) -ᵥ q = g - (q -ᵥ p) := by rw [vadd_vsub_assoc, sub_eq_add_neg, neg_vsub_eq_vsub_rev] /-- Subtracting the result of adding a group element produces the same result @@ -171,7 +171,7 @@ theorem singleton_vsub_self (p : P) : ({p} : Set P) -ᵥ {p} = {(0 : G)} := by end Set @[simp] -theorem vadd_vsub_vadd_cancel_right (v₁ v₂ : G) (p : P) : v₁ +ᵥ p -ᵥ (v₂ +ᵥ p) = v₁ - v₂ := by +theorem vadd_vsub_vadd_cancel_right (v₁ v₂ : G) (p : P) : (v₁ +ᵥ p) -ᵥ (v₂ +ᵥ p) = v₁ - v₂ := by rw [vsub_vadd_eq_vsub_sub, vadd_vsub_assoc, vsub_self, add_zero] /-- If the same point subtracted from two points produces equal @@ -218,10 +218,10 @@ theorem vsub_sub_vsub_cancel_left (p₁ p₂ p₃ : P) : p₃ -ᵥ p₂ - (p₃ rw [sub_eq_add_neg, neg_vsub_eq_vsub_rev, add_comm, vsub_add_vsub_cancel] @[simp] -theorem vadd_vsub_vadd_cancel_left (v : G) (p₁ p₂ : P) : v +ᵥ p₁ -ᵥ (v +ᵥ p₂) = p₁ -ᵥ p₂ := by +theorem vadd_vsub_vadd_cancel_left (v : G) (p₁ p₂ : P) : (v +ᵥ p₁) -ᵥ (v +ᵥ p₂) = p₁ -ᵥ p₂ := by rw [vsub_vadd_eq_vsub_sub, vadd_vsub_assoc, add_sub_cancel_left] -theorem vsub_vadd_comm (p₁ p₂ p₃ : P) : (p₁ -ᵥ p₂ : G) +ᵥ p₃ = p₃ -ᵥ p₂ +ᵥ p₁ := by +theorem vsub_vadd_comm (p₁ p₂ p₃ : P) : (p₁ -ᵥ p₂ : G) +ᵥ p₃ = (p₃ -ᵥ p₂) +ᵥ p₁ := by rw [← @vsub_eq_zero_iff_eq G, vadd_vsub_assoc, vsub_vadd_eq_vsub_sub] simp @@ -372,7 +372,7 @@ open Function def pointReflection (x : P) : Perm P := (constVSub x).trans (vaddConst x) -theorem pointReflection_apply (x y : P) : pointReflection x y = x -ᵥ y +ᵥ x := +theorem pointReflection_apply (x y : P) : pointReflection x y = (x -ᵥ y) +ᵥ x := rfl @[simp] diff --git a/Mathlib/Algebra/Group/Action/Defs.lean b/Mathlib/Algebra/Group/Action/Defs.lean index 294281be04be5..fdd34cf9f6af7 100644 --- a/Mathlib/Algebra/Group/Action/Defs.lean +++ b/Mathlib/Algebra/Group/Action/Defs.lean @@ -65,7 +65,7 @@ class AddAction (G : Type*) (P : Type*) [AddMonoid G] extends VAdd G P where /-- Zero is a neutral element for `+ᵥ` -/ protected zero_vadd : ∀ p : P, (0 : G) +ᵥ p = p /-- Associativity of `+` and `+ᵥ` -/ - add_vadd : ∀ (g₁ g₂ : G) (p : P), g₁ + g₂ +ᵥ p = g₁ +ᵥ (g₂ +ᵥ p) + add_vadd : ∀ (g₁ g₂ : G) (p : P), (g₁ + g₂) +ᵥ p = g₁ +ᵥ g₂ +ᵥ p /-- Typeclass for multiplicative actions by monoids. This generalizes group actions. -/ @[to_additive (attr := ext)] @@ -140,7 +140,7 @@ instance smulCommClass_self (M α : Type*) [CommMonoid M] [MulAction M α] : SMu determined by the additive actions of `M` on `N` and `N` on `α`. -/ class VAddAssocClass (M N α : Type*) [VAdd M N] [VAdd N α] [VAdd M α] : Prop where /-- Associativity of `+ᵥ` -/ - vadd_assoc : ∀ (x : M) (y : N) (z : α), x +ᵥ y +ᵥ z = x +ᵥ (y +ᵥ z) + vadd_assoc : ∀ (x : M) (y : N) (z : α), (x +ᵥ y) +ᵥ z = x +ᵥ y +ᵥ z /-- An instance of `IsScalarTower M N α` states that the multiplicative action of `M` on `α` is determined by the multiplicative actions of `M` on `N` diff --git a/Mathlib/Algebra/Group/Operations.lean b/Mathlib/Algebra/Group/Operations.lean index dcc6957a258a7..74d54a8784207 100644 --- a/Mathlib/Algebra/Group/Operations.lean +++ b/Mathlib/Algebra/Group/Operations.lean @@ -82,7 +82,7 @@ class SMul (M : Type u) (α : Type v) where but it is intended to be used for left actions. -/ smul : M → α → α -@[inherit_doc] infixl:65 " +ᵥ " => HVAdd.hVAdd +@[inherit_doc] infixr:65 " +ᵥ " => HVAdd.hVAdd @[inherit_doc] infixl:65 " -ᵥ " => VSub.vsub @[inherit_doc] infixr:73 " • " => HSMul.hSMul diff --git a/Mathlib/Analysis/Normed/Affine/Isometry.lean b/Mathlib/Analysis/Normed/Affine/Isometry.lean index d6db878a74fdd..ad879f1396054 100644 --- a/Mathlib/Analysis/Normed/Affine/Isometry.lean +++ b/Mathlib/Analysis/Normed/Affine/Isometry.lean @@ -650,7 +650,7 @@ def pointReflection (x : P) : P ≃ᵃⁱ[𝕜] P := variable {𝕜} -theorem pointReflection_apply (x y : P) : (pointReflection 𝕜 x) y = x -ᵥ y +ᵥ x := +theorem pointReflection_apply (x y : P) : (pointReflection 𝕜 x) y = (x -ᵥ y) +ᵥ x := rfl @[simp] diff --git a/Mathlib/Geometry/Euclidean/Basic.lean b/Mathlib/Geometry/Euclidean/Basic.lean index d39b79aad2703..f708b2d5f2073 100644 --- a/Mathlib/Geometry/Euclidean/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Basic.lean @@ -472,14 +472,14 @@ general sense of the word that includes both those common cases. -/ def reflection (s : AffineSubspace ℝ P) [Nonempty s] [HasOrthogonalProjection s.direction] : P ≃ᵃⁱ[ℝ] P := AffineIsometryEquiv.mk' - (fun p => ↑(orthogonalProjection s p) -ᵥ p +ᵥ (orthogonalProjection s p : P)) + (fun p => (↑(orthogonalProjection s p) -ᵥ p) +ᵥ (orthogonalProjection s p : P)) (_root_.reflection s.direction) (↑(Classical.arbitrary s)) (by intro p let v := p -ᵥ ↑(Classical.arbitrary s) let a : V := _root_.orthogonalProjection s.direction v let b : P := ↑(Classical.arbitrary s) - have key : a +ᵥ b -ᵥ (v +ᵥ b) +ᵥ (a +ᵥ b) = a + a - v +ᵥ (b -ᵥ b +ᵥ b) := by + have key : ((a +ᵥ b) -ᵥ (v +ᵥ b)) +ᵥ (a +ᵥ b) = (a + a - v) +ᵥ (b -ᵥ b) +ᵥ b := by rw [← add_vadd, vsub_vadd_eq_vsub_sub, vsub_vadd, vadd_vsub] congr 1 abel @@ -489,7 +489,8 @@ def reflection (s : AffineSubspace ℝ P) [Nonempty s] [HasOrthogonalProjection /-- The result of reflecting. -/ theorem reflection_apply (s : AffineSubspace ℝ P) [Nonempty s] [HasOrthogonalProjection s.direction] - (p : P) : reflection s p = ↑(orthogonalProjection s p) -ᵥ p +ᵥ (orthogonalProjection s p : P) := + (p : P) : + reflection s p = (↑(orthogonalProjection s p) -ᵥ p) +ᵥ (orthogonalProjection s p : P) := rfl theorem eq_reflection_of_eq_subspace {s s' : AffineSubspace ℝ P} [Nonempty s] [Nonempty s'] diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean index 75e11673b7d74..1ac5e28cebf7a 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean @@ -467,7 +467,7 @@ open Function def pointReflection (x : P₁) : P₁ ≃ᵃ[k] P₁ := (constVSub k x).trans (vaddConst k x) -theorem pointReflection_apply (x y : P₁) : pointReflection k x y = x -ᵥ y +ᵥ x := +theorem pointReflection_apply (x y : P₁) : pointReflection k x y = (x -ᵥ y) +ᵥ x := rfl @[simp] diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean index 4ad30ae06cd1d..76c5638c80a15 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean @@ -258,7 +258,8 @@ theorem coe_direction_eq_vsub_set_right {s : AffineSubspace k P} {p : P} (hp : p rw [coe_direction_eq_vsub_set ⟨p, hp⟩] refine le_antisymm ?_ ?_ · rintro v ⟨p1, hp1, p2, hp2, rfl⟩ - exact ⟨p1 -ᵥ p2 +ᵥ p, vadd_mem_of_mem_direction (vsub_mem_direction hp1 hp2) hp, vadd_vsub _ _⟩ + exact ⟨(p1 -ᵥ p2) +ᵥ p, + vadd_mem_of_mem_direction (vsub_mem_direction hp1 hp2) hp, vadd_vsub _ _⟩ · rintro v ⟨p2, hp2, rfl⟩ exact ⟨p2, hp2, p, hp, rfl⟩ @@ -686,7 +687,7 @@ theorem direction_top : (⊤ : AffineSubspace k P).direction = ⊤ := by cases' S.nonempty with p ext v refine ⟨imp_intro Submodule.mem_top, fun _hv => ?_⟩ - have hpv : (v +ᵥ p -ᵥ p : V) ∈ (⊤ : AffineSubspace k P).direction := + have hpv : ((v +ᵥ p) -ᵥ p : V) ∈ (⊤ : AffineSubspace k P).direction := vsub_mem_direction (mem_top k V _) (mem_top k V _) rwa [vadd_vsub] at hpv From 6ce78a15d6de88c5f965e3287cda0359419ef08e Mon Sep 17 00:00:00 2001 From: Yongle Hu Date: Sat, 30 Nov 2024 00:03:03 +0000 Subject: [PATCH 426/829] feat(Algebra/Polynomial/Roots): for a monic polynomial `p`, `p.roots.map f = (p.map f).roots` if `p.roots.card = p.natDegree` (#19430) A variant of the theorem [Polynomial.roots_map_of_injective_of_card_eq_natDegree](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Algebra/Polynomial/Roots.html#Polynomial.roots_map_of_injective_of_card_eq_natDegree) that replaces the injectivity condition with `p.Monic`. Co-authored-by: Yongle Hu <140475041+mbkybky@users.noreply.github.com> Co-authored-by: Yongle Hu <2065545849@qq.com> --- Mathlib/Algebra/Polynomial/Roots.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/Mathlib/Algebra/Polynomial/Roots.lean b/Mathlib/Algebra/Polynomial/Roots.lean index d2a32fe42bc6e..2078ef2aeeb5f 100644 --- a/Mathlib/Algebra/Polynomial/Roots.lean +++ b/Mathlib/Algebra/Polynomial/Roots.lean @@ -751,6 +751,16 @@ theorem roots_map_of_injective_of_card_eq_natDegree [IsDomain A] [IsDomain B] {p apply Multiset.eq_of_le_of_card_le (map_roots_le_of_injective p hf) simpa only [Multiset.card_map, hroots] using (card_roots' _).trans (natDegree_map_le f p) +theorem roots_map_of_map_ne_zero_of_card_eq_natDegree [IsDomain A] [IsDomain B] {p : A[X]} + (f : A →+* B) (h : p.map f ≠ 0) (hroots : p.roots.card = p.natDegree) : + p.roots.map f = (p.map f).roots := + eq_of_le_of_card_le (map_roots_le h) <| by + simpa only [Multiset.card_map, hroots] using (p.map f).card_roots'.trans (p.natDegree_map_le f) + +theorem Monic.roots_map_of_card_eq_natDegree [IsDomain A] [IsDomain B] {p : A[X]} (hm : p.Monic) + (f : A →+* B) (hroots : p.roots.card = p.natDegree) : p.roots.map f = (p.map f).roots := + roots_map_of_map_ne_zero_of_card_eq_natDegree f (map_monic_ne_zero hm) hroots + end end Polynomial From 8a86c2c4063106bdb8e73bb1ac4b9ff5d8366300 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Sat, 30 Nov 2024 00:03:04 +0000 Subject: [PATCH 427/829] feat(Algebra/Homology/Embedding): the canonical truncation functor `trunGEFunctor` (#19543) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Given an embedding `e : Embedding c c'` of complex shapes which satisfy `e.IsTruncGE`, we construct the canonical truncation functor `e.truncGEFunctor C : HomologicalComplex C c' ⥤ HomologicalComplex C c'`. --- .../Algebra/Homology/Embedding/TruncGE.lean | 98 +++++++++++++++++++ 1 file changed, 98 insertions(+) diff --git a/Mathlib/Algebra/Homology/Embedding/TruncGE.lean b/Mathlib/Algebra/Homology/Embedding/TruncGE.lean index c0c9a8b3cbd18..8bac0da93a294 100644 --- a/Mathlib/Algebra/Homology/Embedding/TruncGE.lean +++ b/Mathlib/Algebra/Homology/Embedding/TruncGE.lean @@ -128,6 +128,8 @@ lemma truncGE'_d_eq_fromOpcycles {i j : ι} (hij : c.Rel i j) {i' j' : ι'} subst hi' hj' simp [truncGE'XIso, truncGE'XIsoOpcycles] +section + variable [HasZeroObject C] /-- The canonical truncation of a homological complex relative to an embedding @@ -146,4 +148,100 @@ noncomputable def truncGEXIsoOpcycles {i : ι} {i' : ι'} (hi' : e.f i = i') (hi (K.truncGE e).X i' ≅ K.opcycles i' := (K.truncGE' e).extendXIso e hi' ≪≫ K.truncGE'XIsoOpcycles e hi' hi +end + +section + +variable {K L M} + +open Classical in +/-- The morphism `K.truncGE' e ⟶ L.truncGE' e` induced by a morphism `K ⟶ L`. -/ +noncomputable def truncGE'Map : K.truncGE' e ⟶ L.truncGE' e where + f i := + if hi : e.BoundaryGE i + then + (K.truncGE'XIsoOpcycles e rfl hi).hom ≫ opcyclesMap φ (e.f i) ≫ + (L.truncGE'XIsoOpcycles e rfl hi).inv + else + (K.truncGE'XIso e rfl hi).hom ≫ φ.f (e.f i) ≫ (L.truncGE'XIso e rfl hi).inv + comm' i j hij := by + dsimp + rw [dif_neg (e.not_boundaryGE_next hij)] + by_cases hi : e.BoundaryGE i + · rw [dif_pos hi] + simp [truncGE'_d_eq_fromOpcycles _ e hij rfl rfl hi, + ← cancel_epi (K.pOpcycles (e.f i))] + · rw [dif_neg hi] + simp [truncGE'_d_eq _ e hij rfl rfl hi] + +lemma truncGE'Map_f_eq_opcyclesMap {i : ι} (hi : e.BoundaryGE i) {i' : ι'} (h : e.f i = i') : + (truncGE'Map φ e).f i = + (K.truncGE'XIsoOpcycles e h hi).hom ≫ opcyclesMap φ i' ≫ + (L.truncGE'XIsoOpcycles e h hi).inv := by + subst h + exact dif_pos hi + +lemma truncGE'Map_f_eq {i : ι} (hi : ¬ e.BoundaryGE i) {i' : ι'} (h : e.f i = i') : + (truncGE'Map φ e).f i = + (K.truncGE'XIso e h hi).hom ≫ φ.f i' ≫ (L.truncGE'XIso e h hi).inv := by + subst h + exact dif_neg hi + +variable (K) in +@[simp] +lemma truncGE'Map_id : truncGE'Map (𝟙 K) e = 𝟙 _ := by + ext i + by_cases hi : e.BoundaryGE i + · simp [truncGE'Map_f_eq_opcyclesMap _ _ hi rfl] + · simp [truncGE'Map_f_eq _ _ hi rfl] + +@[reassoc, simp] +lemma truncGE'Map_comp : truncGE'Map (φ ≫ φ') e = truncGE'Map φ e ≫ truncGE'Map φ' e := by + ext i + by_cases hi : e.BoundaryGE i + · simp [truncGE'Map_f_eq_opcyclesMap _ _ hi rfl, opcyclesMap_comp] + · simp [truncGE'Map_f_eq _ _ hi rfl] + +variable [HasZeroObject C] + +/-- The morphism `K.truncGE e ⟶ L.truncGE e` induced by a morphism `K ⟶ L`. -/ +noncomputable def truncGEMap : K.truncGE e ⟶ L.truncGE e := + (e.extendFunctor C).map (truncGE'Map φ e) + +variable (K) in +@[simp] +lemma truncGEMap_id : truncGEMap (𝟙 K) e = 𝟙 _ := by + simp [truncGEMap, truncGE] + +@[reassoc, simp] +lemma truncGEMap_comp : truncGEMap (φ ≫ φ') e = truncGEMap φ e ≫ truncGEMap φ' e := by + simp [truncGEMap, truncGE] + +end + end HomologicalComplex + +namespace ComplexShape.Embedding + +variable (e : Embedding c c') [e.IsTruncGE] + (C : Type*) [Category C] [HasZeroMorphisms C] [HasZeroObject C] [CategoryWithHomology C] + +/-- Given an embedding `e : Embedding c c'` of complex shapes which satisfy `e.IsTruncGE`, +this is the (canonical) truncation functor +`HomologicalComplex C c' ⥤ HomologicalComplex C c`. -/ +@[simps] +noncomputable def truncGE'Functor : + HomologicalComplex C c' ⥤ HomologicalComplex C c where + obj K := K.truncGE' e + map φ := HomologicalComplex.truncGE'Map φ e + +/-- Given an embedding `e : Embedding c c'` of complex shapes which satisfy `e.IsTruncGE`, +this is the (canonical) truncation functor +`HomologicalComplex C c' ⥤ HomologicalComplex C c'`. -/ +@[simps] +noncomputable def truncGEFunctor : + HomologicalComplex C c' ⥤ HomologicalComplex C c' where + obj K := K.truncGE e + map φ := HomologicalComplex.truncGEMap φ e + +end ComplexShape.Embedding From e1064d2ad49b11600b2737fe3fdee1a27ded294f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Sat, 30 Nov 2024 00:03:05 +0000 Subject: [PATCH 428/829] feat(Algebra/Homology): more duality results (#19559) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Given an homological complex `K`, we introduce duality isomorphisms like `cyclesOpIso : K.op.cycles i ≅ op (K.opcycles i)` and study their naturality. --- Mathlib/Algebra/Homology/Opposite.lean | 104 ++++++++++++++++++ .../ShortComplex/HomologicalComplex.lean | 10 ++ .../Homology/ShortComplex/Homology.lean | 14 ++- .../Homology/ShortComplex/RightHomology.lean | 64 +++++++++++ 4 files changed, 191 insertions(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Homology/Opposite.lean b/Mathlib/Algebra/Homology/Opposite.lean index ba39dfe50ab4e..e234cce59912e 100644 --- a/Mathlib/Algebra/Homology/Opposite.lean +++ b/Mathlib/Algebra/Homology/Opposite.lean @@ -213,6 +213,16 @@ instance (K : HomologicalComplex Vᵒᵖ c) (i : ι) [K.HasHomology i] : K.unop.HasHomology i := (inferInstance : (K.sc i).unop.HasHomology) +instance (K : HomologicalComplex V c) (i : ι) [K.HasHomology i] : + ((opFunctor _ _).obj (op K)).HasHomology i := by + dsimp + infer_instance + +instance (K : HomologicalComplex Vᵒᵖ c) (i : ι) [K.HasHomology i] : + ((unopFunctor _ _).obj (op K)).HasHomology i := by + dsimp + infer_instance + variable {V c} /-- If `K` is a homological complex, then the homology of `K.op` identifies to @@ -227,6 +237,100 @@ def homologyUnop (K : HomologicalComplex Vᵒᵖ c) (i : ι) [K.HasHomology i] : K.unop.homology i ≅ unop (K.homology i) := (K.unop.homologyOp i).unop +section + +variable (K : HomologicalComplex V c) (i : ι) [K.HasHomology i] + +/-- The canonical isomorphism `K.op.cycles i ≅ op (K.opcycles i)`. -/ +def cyclesOpIso : K.op.cycles i ≅ op (K.opcycles i) := + (K.sc i).cyclesOpIso + +/-- The canonical isomorphism `K.op.opcycles i ≅ op (K.cycles i)`. -/ +def opcyclesOpIso : K.op.opcycles i ≅ op (K.cycles i) := + (K.sc i).opcyclesOpIso + +variable (j : ι) + +@[reassoc (attr := simp)] +lemma opcyclesOpIso_hom_toCycles_op : + (K.opcyclesOpIso i).hom ≫ (K.toCycles j i).op = K.op.fromOpcycles i j := by + by_cases hij : c.Rel j i + · obtain rfl := c.prev_eq' hij + exact (K.sc i).opcyclesOpIso_hom_toCycles_op + · rw [K.toCycles_eq_zero hij, K.op.fromOpcycles_eq_zero hij, op_zero, comp_zero] + +@[reassoc (attr := simp)] +lemma fromOpcycles_op_cyclesOpIso_inv : + (K.fromOpcycles i j).op ≫ (K.cyclesOpIso i).inv = K.op.toCycles j i := by + by_cases hij : c.Rel i j + · obtain rfl := c.next_eq' hij + exact (K.sc i).fromOpcycles_op_cyclesOpIso_inv + · rw [K.op.toCycles_eq_zero hij, K.fromOpcycles_eq_zero hij, op_zero, zero_comp] + +end + +section + +variable {K L : HomologicalComplex V c} (φ : K ⟶ L) (i : ι) + [K.HasHomology i] [L.HasHomology i] + +@[reassoc] +lemma homologyOp_hom_naturality : + homologyMap ((opFunctor _ _).map φ.op) _ ≫ (K.homologyOp i).hom = + (L.homologyOp i).hom ≫ (homologyMap φ i).op := + ShortComplex.homologyOpIso_hom_naturality ((shortComplexFunctor V c i).map φ) + +@[reassoc] +lemma opcyclesOpIso_hom_naturality : + opcyclesMap ((opFunctor _ _).map φ.op) _ ≫ (K.opcyclesOpIso i).hom = + (L.opcyclesOpIso i).hom ≫ (cyclesMap φ i).op := + ShortComplex.opcyclesOpIso_hom_naturality ((shortComplexFunctor V c i).map φ) + +@[reassoc] +lemma opcyclesOpIso_inv_naturality : + (cyclesMap φ i).op ≫ (K.opcyclesOpIso i).inv = + (L.opcyclesOpIso i).inv ≫ opcyclesMap ((opFunctor _ _).map φ.op) _ := + ShortComplex.opcyclesOpIso_inv_naturality ((shortComplexFunctor V c i).map φ) + +@[reassoc] +lemma cyclesOpIso_hom_naturality : + cyclesMap ((opFunctor _ _).map φ.op) _ ≫ (K.cyclesOpIso i).hom = + (L.cyclesOpIso i).hom ≫ (opcyclesMap φ i).op := + ShortComplex.cyclesOpIso_hom_naturality ((shortComplexFunctor V c i).map φ) + +@[reassoc] +lemma cyclesOpIso_inv_naturality : + (opcyclesMap φ i).op ≫ (K.cyclesOpIso i).inv = + (L.cyclesOpIso i).inv ≫ cyclesMap ((opFunctor _ _).map φ.op) _ := + ShortComplex.cyclesOpIso_inv_naturality ((shortComplexFunctor V c i).map φ) + +end + +section + +variable (V c) [CategoryWithHomology V] (i : ι) + +/-- The natural isomorphism `K.op.cycles i ≅ op (K.opcycles i)`. -/ +@[simps!] +def cyclesOpNatIso : + opFunctor V c ⋙ cyclesFunctor Vᵒᵖ c.symm i ≅ (opcyclesFunctor V c i).op := + NatIso.ofComponents (fun K ↦ (unop K).cyclesOpIso i) + (fun _ ↦ cyclesOpIso_hom_naturality _ _) + +/-- The natural isomorphism `K.op.opcycles i ≅ op (K.cycles i)`. -/ +def opcyclesOpNatIso : + opFunctor V c ⋙ opcyclesFunctor Vᵒᵖ c.symm i ≅ (cyclesFunctor V c i).op := + NatIso.ofComponents (fun K ↦ (unop K).opcyclesOpIso i) + (fun _ ↦ opcyclesOpIso_hom_naturality _ _) + +/-- The natural isomorphism `K.op.homology i ≅ op (K.homology i)`. -/ +def homologyOpNatIso : + opFunctor V c ⋙ homologyFunctor Vᵒᵖ c.symm i ≅ (homologyFunctor V c i).op := + NatIso.ofComponents (fun K ↦ (unop K).homologyOp i) + (fun _ ↦ homologyOp_hom_naturality _ _) + +end + end section diff --git a/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean b/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean index fb3e8617e352d..415cac78ef15f 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean @@ -149,6 +149,11 @@ lemma d_toCycles [K.HasHomology k] : K.d i j ≫ K.toCycles j k = 0 := by simp only [← cancel_mono (K.iCycles k), assoc, toCycles_i, d_comp_d, zero_comp] +variable {i j} in +lemma toCycles_eq_zero [K.HasHomology j] (hij : ¬ c.Rel i j) : + K.toCycles i j = 0 := by + rw [← cancel_mono (K.iCycles j), toCycles_i, zero_comp, K.shape _ _ hij] + variable {i} section @@ -258,6 +263,11 @@ lemma fromOpcycles_d : K.fromOpcycles i j ≫ K.d j k = 0 := by simp only [← cancel_epi (K.pOpcycles i), p_fromOpcycles_assoc, d_comp_d, comp_zero] +variable {i j} in +lemma fromOpcycles_eq_zero (hij : ¬ c.Rel i j) : + K.fromOpcycles i j = 0 := by + rw [← cancel_epi (K.pOpcycles i), p_fromOpcycles, comp_zero, K.shape _ _ hij] + variable {i} @[reassoc] diff --git a/Mathlib/Algebra/Homology/ShortComplex/Homology.lean b/Mathlib/Algebra/Homology/ShortComplex/Homology.lean index 0c416a49f4df5..9f3d4faff0bed 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/Homology.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/Homology.lean @@ -1036,6 +1036,18 @@ lemma homologyMap_op [HasHomology S₁] [HasHomology S₂] : simp only [assoc, rightHomologyMap'_op, op_comp, ← leftHomologyMap'_comp_assoc, id_comp, opMap_id, comp_id, HomologyData.op_left] +@[reassoc] +lemma homologyOpIso_hom_naturality [S₁.HasHomology] [S₂.HasHomology] : + homologyMap (opMap φ) ≫ (S₁.homologyOpIso).hom = + S₂.homologyOpIso.hom ≫ (homologyMap φ).op := by + simp [homologyMap_op] + +@[reassoc] +lemma homologyOpIso_inv_naturality [S₁.HasHomology] [S₂.HasHomology] : + (homologyMap φ).op ≫ (S₁.homologyOpIso).inv = + S₂.homologyOpIso.inv ≫ homologyMap (opMap φ) := by + simp [homologyMap_op] + variable (C) /-- The natural isomorphism `(homologyFunctor C).op ≅ opFunctor C ⋙ homologyFunctor Cᵒᵖ` @@ -1043,7 +1055,7 @@ which relates the homology in `C` and in `Cᵒᵖ`. -/ noncomputable def homologyFunctorOpNatIso [CategoryWithHomology C] : (homologyFunctor C).op ≅ opFunctor C ⋙ homologyFunctor Cᵒᵖ := NatIso.ofComponents (fun S => S.unop.homologyOpIso.symm) - (by simp [homologyMap_op]) + (fun _ ↦ homologyOpIso_inv_naturality _) variable {C} {A : C} diff --git a/Mathlib/Algebra/Homology/ShortComplex/RightHomology.lean b/Mathlib/Algebra/Homology/ShortComplex/RightHomology.lean index 3592289533893..9832970621a74 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/RightHomology.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/RightHomology.lean @@ -964,6 +964,70 @@ noncomputable def cyclesOpIso [S.HasRightHomology] : S.op.cycles ≅ Opposite.op S.opcycles := S.rightHomologyData.op.cyclesIso +@[reassoc (attr := simp)] +lemma opcyclesOpIso_hom_toCycles_op [S.HasLeftHomology] : + S.opcyclesOpIso.hom ≫ S.toCycles.op = S.op.fromOpcycles := by + dsimp [opcyclesOpIso, toCycles] + rw [← cancel_epi S.op.pOpcycles, p_fromOpcycles, + RightHomologyData.pOpcycles_comp_opcyclesIso_hom_assoc, + LeftHomologyData.op_p, ← op_comp, LeftHomologyData.f'_i, op_g] + +@[reassoc (attr := simp)] +lemma fromOpcycles_op_cyclesOpIso_inv [S.HasRightHomology]: + S.fromOpcycles.op ≫ S.cyclesOpIso.inv = S.op.toCycles := by + dsimp [cyclesOpIso, fromOpcycles] + rw [← cancel_mono S.op.iCycles, assoc, toCycles_i, + LeftHomologyData.cyclesIso_inv_comp_iCycles, RightHomologyData.op_i, + ← op_comp, RightHomologyData.p_g', op_f] + +@[reassoc (attr := simp)] +lemma op_pOpcycles_opcyclesOpIso_hom [S.HasLeftHomology] : + S.op.pOpcycles ≫ S.opcyclesOpIso.hom = S.iCycles.op := by + dsimp [opcyclesOpIso] + rw [← S.leftHomologyData.op.p_comp_opcyclesIso_inv, assoc, + Iso.inv_hom_id, comp_id] + rfl + +@[reassoc (attr := simp)] +lemma cyclesOpIso_inv_op_iCycles [S.HasRightHomology] : + S.cyclesOpIso.inv ≫ S.op.iCycles = S.pOpcycles.op := by + dsimp [cyclesOpIso] + rw [← S.rightHomologyData.op.cyclesIso_hom_comp_i, Iso.inv_hom_id_assoc] + rfl + +@[reassoc] +lemma opcyclesOpIso_hom_naturality (φ : S₁ ⟶ S₂) + [S₁.HasLeftHomology] [S₂.HasLeftHomology] : + opcyclesMap (opMap φ) ≫ (S₁.opcyclesOpIso).hom = + S₂.opcyclesOpIso.hom ≫ (cyclesMap φ).op := by + rw [← cancel_epi S₂.op.pOpcycles, p_opcyclesMap_assoc, opMap_τ₂, + op_pOpcycles_opcyclesOpIso_hom, op_pOpcycles_opcyclesOpIso_hom_assoc, ← op_comp, + ← op_comp, cyclesMap_i] + +@[reassoc] +lemma opcyclesOpIso_inv_naturality (φ : S₁ ⟶ S₂) + [S₁.HasLeftHomology] [S₂.HasLeftHomology] : + (cyclesMap φ).op ≫ (S₁.opcyclesOpIso).inv = + S₂.opcyclesOpIso.inv ≫ opcyclesMap (opMap φ) := by + rw [← cancel_epi (S₂.opcyclesOpIso.hom), Iso.hom_inv_id_assoc, + ← opcyclesOpIso_hom_naturality_assoc, Iso.hom_inv_id, comp_id] + +@[reassoc] +lemma cyclesOpIso_inv_naturality (φ : S₁ ⟶ S₂) + [S₁.HasRightHomology] [S₂.HasRightHomology] : + (opcyclesMap φ).op ≫ (S₁.cyclesOpIso).inv = + S₂.cyclesOpIso.inv ≫ cyclesMap (opMap φ) := by + rw [← cancel_mono S₁.op.iCycles, assoc, assoc, cyclesOpIso_inv_op_iCycles, cyclesMap_i, + cyclesOpIso_inv_op_iCycles_assoc, ← op_comp, p_opcyclesMap, op_comp, opMap_τ₂] + +@[reassoc] +lemma cyclesOpIso_hom_naturality (φ : S₁ ⟶ S₂) + [S₁.HasRightHomology] [S₂.HasRightHomology] : + cyclesMap (opMap φ) ≫ (S₁.cyclesOpIso).hom = + S₂.cyclesOpIso.hom ≫ (opcyclesMap φ).op := by + rw [← cancel_mono (S₁.cyclesOpIso).inv, assoc, assoc, Iso.hom_inv_id, comp_id, + cyclesOpIso_inv_naturality, Iso.hom_inv_id_assoc] + @[simp] lemma leftHomologyMap'_op (φ : S₁ ⟶ S₂) (h₁ : S₁.LeftHomologyData) (h₂ : S₂.LeftHomologyData) : From 21608d1f823014e49969e44316d30c5ea6de2335 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Sat, 30 Nov 2024 00:03:07 +0000 Subject: [PATCH 429/829] feat(Algebra/Homology/Embedding): dualise extend (#19560) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit When `e` is an embedding of complex shapes, we obtain an isomorphism `extendOpIso : K.op.extend e.op ≅ (K.extend e).op`. --- .../Algebra/Homology/Embedding/Extend.lean | 36 +++++++++++++++++++ 1 file changed, 36 insertions(+) diff --git a/Mathlib/Algebra/Homology/Embedding/Extend.lean b/Mathlib/Algebra/Homology/Embedding/Extend.lean index 5e2f3e79fe10c..6233b7857282a 100644 --- a/Mathlib/Algebra/Homology/Embedding/Extend.lean +++ b/Mathlib/Algebra/Homology/Embedding/Extend.lean @@ -5,6 +5,7 @@ Authors: Joël Riou -/ import Mathlib.Algebra.Homology.Embedding.Basic import Mathlib.Algebra.Homology.Additive +import Mathlib.Algebra.Homology.Opposite /-! # The extension of a homological complex by an embedding of complex shapes @@ -46,6 +47,12 @@ lemma isZero_X {i : Option ι} (hi : i = none) : subst hi exact Limits.isZero_zero _ +/-- The canonical isomorphism `X K.op i ≅ Opposite.op (X K i)`. -/ +noncomputable def XOpIso (i : Option ι) : X K.op i ≅ Opposite.op (X K i) := + match i with + | some _ => Iso.refl _ + | none => IsZero.iso (isZero_X _ rfl) (isZero_X K rfl).op + /-- Auxiliary definition for the `d` field of `HomologicalComplex.extend`. -/ noncomputable def d : ∀ (i j : Option ι), extend.X K i ⟶ extend.X K j | none, _ => 0 @@ -64,6 +71,21 @@ lemma d_eq {i j : Option ι} {a b : ι} (hi : i = some a) (hj : j = some b) : dsimp [XIso, d] erw [id_comp, comp_id] +@[reassoc] +lemma XOpIso_hom_d_op (i j : Option ι) : + (XOpIso K i).hom ≫ (d K j i).op = + d K.op i j ≫ (XOpIso K j).hom := + match i, j with + | none, _ => by + simp only [d_none_eq_zero, d_none_eq_zero', comp_zero, zero_comp, op_zero] + | some i, some j => by + dsimp [XOpIso] + simp only [d_eq _ rfl rfl, Option.some.injEq, d_eq, op_comp, assoc, + id_comp, comp_id] + rfl + | some _, none => by + simp only [d_none_eq_zero, d_none_eq_zero', comp_zero, zero_comp, op_zero] + variable {K L} /-- Auxiliary definition for `HomologicalComplex.extendMap`. -/ @@ -222,6 +244,20 @@ lemma extendMap_zero : extendMap (0 : K ⟶ L) e = 0 := by simp [extendMap_f _ e hi] · apply (K.isZero_extend_X e i' (fun i hi => hi' ⟨i, hi⟩)).eq_of_src +/-- The canonical isomorphism `K.op.extend e.op ≅ (K.extend e).op`. -/ +noncomputable def extendOpIso : K.op.extend e.op ≅ (K.extend e).op := + Hom.isoOfComponents (fun _ ↦ extend.XOpIso _ _) (fun _ _ _ ↦ + extend.XOpIso_hom_d_op _ _ _) + +@[reassoc] +lemma extend_op_d (i' j' : ι') : + (K.op.extend e.op).d i' j' = + (K.extendOpIso e).hom.f i' ≫ ((K.extend e).d j' i').op ≫ + (K.extendOpIso e).inv.f j' := by + have := (K.extendOpIso e).inv.comm i' j' + dsimp at this + rw [← this, ← comp_f_assoc, Iso.hom_inv_id, id_f, id_comp] + end @[simp] From eb9d9b39a42ac6b69b6f2c56de429cb31ac80e69 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Sat, 30 Nov 2024 01:55:31 +0000 Subject: [PATCH 430/829] chore: fix names `{un}op_non{neg,pos}` (#19575) --- Mathlib/Algebra/Order/Group/Opposite.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/Algebra/Order/Group/Opposite.lean b/Mathlib/Algebra/Order/Group/Opposite.lean index e145f616c4c11..bf68704269209 100644 --- a/Mathlib/Algebra/Order/Group/Opposite.lean +++ b/Mathlib/Algebra/Order/Group/Opposite.lean @@ -51,10 +51,10 @@ 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 +@[simp] lemma unop_nonpos {a : αᵐᵒᵖ} : unop a ≤ 0 ↔ a ≤ 0 := .rfl +@[simp] lemma unop_nonneg {a : αᵐᵒᵖ} : 0 ≤ unop a ↔ 0 ≤ a := .rfl +@[simp] lemma op_nonpos {a : α} : op a ≤ 0 ↔ a ≤ 0 := .rfl +@[simp] lemma op_nonneg {a : α} : 0 ≤ op a ↔ 0 ≤ a := .rfl end OrderedAddCommMonoid From c73ab4cda826a2c4a78a21bf064c41c4f74b91fb Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Sat, 30 Nov 2024 09:05:29 +0000 Subject: [PATCH 431/829] feat(CategoryTheory/SmallObject/Iteration): existence of objects (bot and succ cases) (#19047) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- Mathlib.lean | 2 + .../SmallObject/Iteration/ExtendToSucc.lean | 161 ++++++++++++++++++ .../SmallObject/Iteration/Nonempty.lean | 97 +++++++++++ 3 files changed, 260 insertions(+) create mode 100644 Mathlib/CategoryTheory/SmallObject/Iteration/ExtendToSucc.lean create mode 100644 Mathlib/CategoryTheory/SmallObject/Iteration/Nonempty.lean diff --git a/Mathlib.lean b/Mathlib.lean index 85bd8c077ab33..2c60994a05392 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2128,6 +2128,8 @@ import Mathlib.CategoryTheory.Sites.Whiskering import Mathlib.CategoryTheory.Skeletal import Mathlib.CategoryTheory.SmallObject.Construction import Mathlib.CategoryTheory.SmallObject.Iteration.Basic +import Mathlib.CategoryTheory.SmallObject.Iteration.ExtendToSucc +import Mathlib.CategoryTheory.SmallObject.Iteration.Nonempty import Mathlib.CategoryTheory.SmallObject.Iteration.UniqueHom import Mathlib.CategoryTheory.Square import Mathlib.CategoryTheory.Subobject.Basic diff --git a/Mathlib/CategoryTheory/SmallObject/Iteration/ExtendToSucc.lean b/Mathlib/CategoryTheory/SmallObject/Iteration/ExtendToSucc.lean new file mode 100644 index 0000000000000..166469022df52 --- /dev/null +++ b/Mathlib/CategoryTheory/SmallObject/Iteration/ExtendToSucc.lean @@ -0,0 +1,161 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import Mathlib.CategoryTheory.SmallObject.Iteration.Basic + +/-! +# Extension of a functor from `Set.Iic j` to `Set.Iic (Order.succ j)` + +Given a linearly ordered type `J` with `SuccOrder J`, `j : J` that is not maximal, +we define the extension of a functor `F : Set.Iic j ⥤ C` as a +functor `Set.Iic (Order.succ j) ⥤ C` when an object `X : C` and a morphism +`τ : F.obj ⟨j, _⟩ ⟶ X` is given. + +-/ + +universe u + +namespace CategoryTheory + +open Category + +namespace Functor + +variable {C : Type*} [Category C] + {J : Type u} [LinearOrder J] [SuccOrder J] {j : J} (hj : ¬IsMax j) + (F : Set.Iic j ⥤ C) {X : C} (τ : F.obj ⟨j, by simp⟩ ⟶ X) + +namespace extendToSucc + +variable (X) + +/-- `extendToSucc`, on objects: it coincides with `F.obj` for `i ≤ j`, and +it sends `Order.succ j` to the given object `X`. -/ +def obj (i : Set.Iic (Order.succ j)) : C := + if hij : i.1 ≤ j then F.obj ⟨i.1, hij⟩ else X + +/-- The isomorphism `obj F X ⟨i, _⟩ ≅ F.obj i` when `i : Set.Iic j`. -/ +def objIso (i : Set.Iic j) : + obj F X ⟨i, i.2.trans (Order.le_succ j)⟩ ≅ F.obj i := eqToIso (dif_pos i.2) + +/-- The isomorphism `obj F X ⟨Order.succ j, _⟩ ≅ X`. -/ +def objSuccIso : + obj F X ⟨Order.succ j, by simp⟩ ≅ X := + eqToIso (dif_neg (by simpa only [Order.succ_le_iff_isMax] using hj)) + +variable {X} + +/-- `extendToSucc`, on morphisms. -/ +def map (i₁ i₂ : J) (hi : i₁ ≤ i₂) (hi₂ : i₂ ≤ Order.succ j) : + obj F X ⟨i₁, hi.trans hi₂⟩ ⟶ obj F X ⟨i₂, hi₂⟩ := + if h₁ : i₂ ≤ j then + (objIso F X ⟨i₁, hi.trans h₁⟩).hom ≫ F.map (homOfLE hi) ≫ (objIso F X ⟨i₂, h₁⟩).inv + else + if h₂ : i₁ ≤ j then + (objIso F X ⟨i₁, h₂⟩).hom ≫ F.map (homOfLE h₂) ≫ τ ≫ + (objSuccIso hj F X).inv ≫ eqToHom (by + congr + exact le_antisymm (Order.succ_le_of_lt (not_le.1 h₁)) hi₂) + else + eqToHom (by + congr + rw [le_antisymm hi₂ (Order.succ_le_of_lt (not_le.1 h₁)), + le_antisymm (hi.trans hi₂) (Order.succ_le_of_lt (not_le.1 h₂))]) + +lemma map_eq (i₁ i₂ : J) (hi : i₁ ≤ i₂) (hi₂ : i₂ ≤ j) : + map hj F τ i₁ i₂ hi (hi₂.trans (Order.le_succ j)) = + (objIso F X ⟨i₁, hi.trans hi₂⟩).hom ≫ F.map (homOfLE hi) ≫ + (objIso F X ⟨i₂, hi₂⟩).inv := + dif_pos hi₂ + +lemma map_self_succ : + map hj F τ j (Order.succ j) (Order.le_succ j) (by rfl) = + (objIso F X ⟨j, by simp⟩).hom ≫ τ ≫ (objSuccIso hj F X).inv := by + dsimp [map] + rw [dif_neg (by simpa only [Order.succ_le_iff_isMax] using hj), + dif_pos (by rfl), map_id, comp_id, id_comp] + +@[simp] +lemma map_id (i : J) (hi : i ≤ Order.succ j) : + map hj F τ i i (by rfl) hi = 𝟙 _ := by + dsimp [map] + by_cases h₁ : i ≤ j + · rw [dif_pos h₁, CategoryTheory.Functor.map_id, id_comp, Iso.hom_inv_id] + · obtain rfl : i = Order.succ j := le_antisymm hi (Order.succ_le_of_lt (not_le.1 h₁)) + rw [dif_neg (by simpa only [Order.succ_le_iff_isMax] using hj), + dif_neg h₁] + +lemma map_comp (i₁ i₂ i₃ : J) (h₁₂ : i₁ ≤ i₂) (h₂₃ : i₂ ≤ i₃) (h : i₃ ≤ Order.succ j) : + map hj F τ i₁ i₃ (h₁₂.trans h₂₃) h = + map hj F τ i₁ i₂ h₁₂ (h₂₃.trans h) ≫ map hj F τ i₂ i₃ h₂₃ h := by + by_cases h₁ : i₃ ≤ j + · rw [map_eq hj F τ i₁ i₂ _ (h₂₃.trans h₁), map_eq hj F τ i₂ i₃ _ h₁, + map_eq hj F τ i₁ i₃ _ h₁, assoc, assoc, Iso.inv_hom_id_assoc, ← map_comp_assoc, + homOfLE_comp] + · obtain rfl : i₃ = Order.succ j := le_antisymm h (Order.succ_le_of_lt (not_le.1 h₁)) + obtain h₂ | rfl := h₂₃.lt_or_eq + · rw [Order.lt_succ_iff_of_not_isMax hj] at h₂ + rw [map_eq hj F τ i₁ i₂ _ h₂] + dsimp [map] + rw [dif_neg h₁, dif_pos (h₁₂.trans h₂), dif_neg h₁, dif_pos h₂, + assoc, assoc, Iso.inv_hom_id_assoc,comp_id, ← map_comp_assoc, homOfLE_comp] + · rw [map_id, comp_id] + +end extendToSucc + +open extendToSucc in +include hj in +/-- The extension to `Set.Iic (Order.succ j) ⥤ C` of a functor `F : Set.Iic j ⥤ C`, +when we specify a morphism `F.obj ⟨j, _⟩ ⟶ X`. -/ +def extendToSucc : Set.Iic (Order.succ j) ⥤ C where + obj := obj F X + map {i₁ i₂} f := map hj F τ i₁ i₂ (leOfHom f) i₂.2 + map_id _ := extendToSucc.map_id _ F τ _ _ + map_comp {i₁ i₂ i₃} f g := extendToSucc.map_comp hj F τ i₁ i₂ i₃ (leOfHom f) (leOfHom g) i₃.2 + +/-- The isomorphism `(extendToSucc hj F τ).obj ⟨i, _⟩ ≅ F.obj i` when `i : Set.Iic j` -/ +def extendToSuccObjIso (i : Set.Iic j) : + (extendToSucc hj F τ).obj ⟨i, i.2.trans (Order.le_succ j)⟩ ≅ F.obj i := + extendToSucc.objIso F X i + +/-- The isomorphism `(extendToSucc hj F τ).obj ⟨Order.succ j, _⟩ ≅ X`. -/ +def extendToSuccObjSuccIso : + (extendToSucc hj F τ).obj ⟨Order.succ j, by simp⟩ ≅ X := + extendToSucc.objSuccIso hj F X + +@[reassoc] +lemma extendToSuccObjIso_hom_naturality (i₁ i₂ : J) (hi : i₁ ≤ i₂) (hi₂ : i₂ ≤ j) : + (extendToSucc hj F τ).map (homOfLE hi : + ⟨i₁, hi.trans (hi₂.trans (Order.le_succ j))⟩ ⟶ ⟨i₂, hi₂.trans (Order.le_succ j)⟩) ≫ + (extendToSuccObjIso hj F τ ⟨i₂, hi₂⟩).hom = + (extendToSuccObjIso hj F τ ⟨i₁, hi.trans hi₂⟩).hom ≫ F.map (homOfLE hi) := by + dsimp [extendToSucc, extendToSuccObjIso] + rw [extendToSucc.map_eq _ _ _ _ _ _ hi₂, assoc, assoc, Iso.inv_hom_id, comp_id] + +/-- The isomorphism expressing that `extendToSucc hj F τ` extends `F`. -/ +@[simps!] +def extendToSuccRestrictionLEIso : + Iteration.restrictionLE (extendToSucc hj F τ) (Order.le_succ j) ≅ F := + NatIso.ofComponents (extendToSuccObjIso hj F τ) (by + rintro ⟨i₁, h₁⟩ ⟨i₂, h₂⟩ f + apply extendToSuccObjIso_hom_naturality) + +lemma extentToSucc_map (i₁ i₂ : J) (hi : i₁ ≤ i₂) (hi₂ : i₂ ≤ j) : + (extendToSucc hj F τ).map (homOfLE hi : + ⟨i₁, hi.trans (hi₂.trans (Order.le_succ j))⟩ ⟶ ⟨i₂, hi₂.trans (Order.le_succ j)⟩) = + (extendToSuccObjIso hj F τ ⟨i₁, hi.trans hi₂⟩).hom ≫ F.map (homOfLE hi) ≫ + (extendToSuccObjIso hj F τ ⟨i₂, hi₂⟩).inv := by + rw [← extendToSuccObjIso_hom_naturality_assoc, Iso.hom_inv_id, comp_id] + +lemma extendToSucc_map_le_succ : + (extendToSucc hj F τ).map (homOfLE (Order.le_succ j)) = + (extendToSuccObjIso hj F τ ⟨j, by simp⟩).hom ≫ τ ≫ + (extendToSuccObjSuccIso hj F τ).inv := + extendToSucc.map_self_succ _ _ _ + +end Functor + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/SmallObject/Iteration/Nonempty.lean b/Mathlib/CategoryTheory/SmallObject/Iteration/Nonempty.lean new file mode 100644 index 0000000000000..80bdbaec34cea --- /dev/null +++ b/Mathlib/CategoryTheory/SmallObject/Iteration/Nonempty.lean @@ -0,0 +1,97 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.CategoryTheory.SmallObject.Iteration.ExtendToSucc + +/-! +# Existence of objects in the category of iterations of functors + +Given a functor `Φ : C ⥤ C` and a natural transformation `ε : 𝟭 C ⟶ Φ`, +we shall show in this file that for any well ordered set `J`, +and `j : J`, the category `Functor.Iteration ε j` is nonempty. +As we already know from the main result in `SmallObject.Iteration.UniqueHom` +that such objects, if they exist, are unique up to a unique isomorphism, +we shall show the existence of a term in `Functor.Iteration ε j` by +transfinite induction. + +-/ + +universe u + +namespace CategoryTheory + +open Category Limits + +variable {C : Type*} [Category C] {Φ : C ⥤ C} {ε : 𝟭 C ⟶ Φ} + {J : Type u} [LinearOrder J] [OrderBot J] [SuccOrder J] + +namespace Functor + +namespace Iteration + +variable (ε J) in +/-- The obvious term in `Iteration ε ⊥`: it is given by the identity functor. -/ +def mkOfBot : Iteration ε (⊥ : J) where + F := (Functor.const _).obj (𝟭 C) + isoZero := Iso.refl _ + isoSucc _ h := by simp at h + mapSucc'_eq _ h := by simp at h + isColimit x hx h := by + exfalso + refine hx.not_isMin (by simpa using h) + +/-- When `j : J` is not maximal, this is the extension as `Iteration ε (Order.succ j)` +of any `iter : Iteration ε j`. -/ +noncomputable def mkOfSucc {j : J} (hj : ¬IsMax j) (iter : Iteration ε j) : + Iteration ε (Order.succ j) where + F := extendToSucc hj iter.F (whiskerLeft _ ε) + isoZero := (extendToSuccObjIso hj iter.F (whiskerLeft _ ε) ⟨⊥, by simp⟩).trans iter.isoZero + isoSucc i hi := + if hij : i < j then + extendToSuccObjIso _ _ _ ⟨Order.succ i, Order.succ_le_of_lt hij⟩ ≪≫ + iter.isoSucc i hij ≪≫ (isoWhiskerRight (extendToSuccObjIso _ _ _ ⟨i, hij.le⟩).symm _) + else + have hij' : i = j := le_antisymm + (by simpa only [Order.lt_succ_iff_of_not_isMax hj] using hi) (by simpa using hij) + eqToIso (by subst hij'; rfl) ≪≫ extendToSuccObjSuccIso hj iter.F (whiskerLeft _ ε) ≪≫ + isoWhiskerRight ((extendToSuccObjIso hj iter.F (whiskerLeft _ ε) ⟨j, by simp⟩).symm.trans + (eqToIso (by subst hij'; rfl))) _ + mapSucc'_eq i hi := by + obtain hi' | rfl := ((Order.lt_succ_iff_of_not_isMax hj).mp hi).lt_or_eq + · ext X + have := iter.mapSucc_eq i hi' + dsimp [mapSucc, mapSucc'] at this ⊢ + rw [extentToSucc_map _ _ _ _ _ _ (Order.succ_le_of_lt hi'), this, dif_pos hi'] + dsimp + rw [assoc, assoc] + erw [ε.naturality_assoc] + · ext X + dsimp [mapSucc'] + rw [dif_neg (gt_irrefl i), extendToSucc_map_le_succ] + dsimp + rw [id_comp, comp_id] + erw [ε.naturality_assoc] + isColimit i hi hij := by + have hij' : i ≤ j := by + obtain hij | rfl := hij.lt_or_eq + · exact (Order.lt_succ_iff_of_not_isMax hj).1 hij + · exfalso + exact Order.not_isSuccLimit_succ_of_not_isMax hj hi + refine (IsColimit.precomposeHomEquiv + (isoWhiskerLeft (monotone_inclusion_lt_le_of_le hij').functor + (extendToSuccRestrictionLEIso hj iter.F (whiskerLeft _ ε))).symm _).1 + (IsColimit.ofIsoColimit (iter.isColimit i hi hij') + (Iso.symm (Cocones.ext (extendToSuccObjIso hj iter.F (whiskerLeft _ ε) ⟨i, hij'⟩) + (fun ⟨k, hk⟩ ↦ ?_)))) + dsimp + rw [assoc, extendToSuccObjIso_hom_naturality hj iter.F (whiskerLeft _ ε)] + dsimp + rw [Iso.inv_hom_id_assoc] + +end Iteration + +end Functor + +end CategoryTheory From ad2e8f1176e800a5d3e6b9083a11601ecf3994ee Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Sat, 30 Nov 2024 09:21:43 +0000 Subject: [PATCH 432/829] chore(RingTheory): split up `Algebraic.lean` (#19370) Split up the file `RingTheory/Algebraic.lean` into a few subfiles: * `Algebraic/Defs.lean`: definition of algebraic and transcendental elements and (sub)algebras * `Algebraic/Integral.lean`: relate `IsAlgebraic` and `IsIntegral` * `Algebraic/Basic.lean`: most results on showing certain elements are algebraic (this is probably the file you want to import afterwards) * `Algebraic/LinearIndependent.lean`: showing certain families are linear independent using transcendentality * `Algebraic/MvPolynomial.lean`: results on multivariate polynomials * `Algebraic/Pi.lean`: a definition of algebraic function (which seems to be unused otherwise). See also [mathlib3#11156](https://github.com/leanprover-community/mathlib3/pull/11156). I don't expect a particularly exciting import reduction, but it does clean up the file structure somewhat, in particular by moving the relation between algebraic and integral elements to one dedicated place (which had been irritating me for years!). Co-authored-by: Anne Baanen --- Mathlib.lean | 7 +- Mathlib/Algebra/AlgebraicCard.lean | 2 +- Mathlib/Algebra/MvPolynomial/Equiv.lean | 5 + Mathlib/Data/Real/GoldenRatio.lean | 7 +- Mathlib/Data/Real/Irrational.lean | 4 +- Mathlib/FieldTheory/AxGrothendieck.lean | 2 +- .../IntermediateField/Algebraic.lean | 4 +- Mathlib/FieldTheory/Minpoly/Field.lean | 3 +- Mathlib/FieldTheory/MvRatFunc/Rank.lean | 3 +- .../SplittingField/Construction.lean | 1 + .../DiophantineApproximation.lean | 1 + .../{Algebraic.lean => Algebraic/Basic.lean} | 363 +----------------- Mathlib/RingTheory/Algebraic/Cardinality.lean | 3 +- Mathlib/RingTheory/Algebraic/Defs.lean | 103 +++++ Mathlib/RingTheory/Algebraic/Integral.lean | 161 ++++++++ .../Algebraic/LinearIndependent.lean | 50 +++ .../RingTheory/Algebraic/MvPolynomial.lean | 101 +++++ Mathlib/RingTheory/Algebraic/Pi.lean | 90 +++++ Mathlib/RingTheory/AlgebraicIndependent.lean | 1 + Mathlib/RingTheory/Localization/Integral.lean | 2 +- MathlibTest/instance_diamonds.lean | 3 +- 21 files changed, 543 insertions(+), 373 deletions(-) rename Mathlib/RingTheory/{Algebraic.lean => Algebraic/Basic.lean} (61%) create mode 100644 Mathlib/RingTheory/Algebraic/Defs.lean create mode 100644 Mathlib/RingTheory/Algebraic/Integral.lean create mode 100644 Mathlib/RingTheory/Algebraic/LinearIndependent.lean create mode 100644 Mathlib/RingTheory/Algebraic/MvPolynomial.lean create mode 100644 Mathlib/RingTheory/Algebraic/Pi.lean diff --git a/Mathlib.lean b/Mathlib.lean index 2c60994a05392..cdccddc0c431c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4214,8 +4214,13 @@ import Mathlib.RingTheory.Adjoin.PowerBasis import Mathlib.RingTheory.Adjoin.Tower import Mathlib.RingTheory.AdjoinRoot import Mathlib.RingTheory.AlgebraTower -import Mathlib.RingTheory.Algebraic +import Mathlib.RingTheory.Algebraic.Basic import Mathlib.RingTheory.Algebraic.Cardinality +import Mathlib.RingTheory.Algebraic.Defs +import Mathlib.RingTheory.Algebraic.Integral +import Mathlib.RingTheory.Algebraic.LinearIndependent +import Mathlib.RingTheory.Algebraic.MvPolynomial +import Mathlib.RingTheory.Algebraic.Pi import Mathlib.RingTheory.AlgebraicIndependent import Mathlib.RingTheory.Artinian import Mathlib.RingTheory.Bezout diff --git a/Mathlib/Algebra/AlgebraicCard.lean b/Mathlib/Algebra/AlgebraicCard.lean index 22df4f242969c..c9fce882650c0 100644 --- a/Mathlib/Algebra/AlgebraicCard.lean +++ b/Mathlib/Algebra/AlgebraicCard.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Violeta Hernández Palacios -/ import Mathlib.Algebra.Polynomial.Cardinal -import Mathlib.RingTheory.Algebraic +import Mathlib.RingTheory.Algebraic.Basic /-! ### Cardinality of algebraic numbers diff --git a/Mathlib/Algebra/MvPolynomial/Equiv.lean b/Mathlib/Algebra/MvPolynomial/Equiv.lean index ebcd5678aa4a4..5080eada19bc1 100644 --- a/Mathlib/Algebra/MvPolynomial/Equiv.lean +++ b/Mathlib/Algebra/MvPolynomial/Equiv.lean @@ -546,6 +546,11 @@ lemma finSuccEquiv_rename_finSuccEquiv (e : σ ≃ Fin n) (φ : MvPolynomial (Op end +@[simp] +theorem rename_polynomial_aeval_X {σ τ : Type*} (f : σ → τ) (i : σ) (p : R[X]) : + rename f (Polynomial.aeval (X i) p) = Polynomial.aeval (X (f i) : MvPolynomial τ R) p := by + rw [← aeval_algHom_apply, rename_X] + end Equiv end MvPolynomial diff --git a/Mathlib/Data/Real/GoldenRatio.lean b/Mathlib/Data/Real/GoldenRatio.lean index 2518fadf3188a..dc7a2c207b503 100644 --- a/Mathlib/Data/Real/GoldenRatio.lean +++ b/Mathlib/Data/Real/GoldenRatio.lean @@ -3,10 +3,11 @@ Copyright (c) 2020 Anatole Dedecker. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anatole Dedecker, Alexey Soloyev, Junyan Xu, Kamila Szewczyk -/ -import Mathlib.Data.Real.Irrational -import Mathlib.Data.Nat.Fib.Basic -import Mathlib.Data.Fin.VecNotation +import Mathlib.Algebra.EuclideanDomain.Basic import Mathlib.Algebra.LinearRecurrence +import Mathlib.Data.Fin.VecNotation +import Mathlib.Data.Nat.Fib.Basic +import Mathlib.Data.Real.Irrational import Mathlib.Tactic.NormNum.NatFib import Mathlib.Tactic.NormNum.Prime diff --git a/Mathlib/Data/Real/Irrational.lean b/Mathlib/Data/Real/Irrational.lean index 59bde31ad0787..235dc1f4d78d5 100644 --- a/Mathlib/Data/Real/Irrational.lean +++ b/Mathlib/Data/Real/Irrational.lean @@ -4,10 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne, Yury Kudryashov -/ import Mathlib.Algebra.Algebra.Rat +import Mathlib.Data.Nat.Prime.Int import Mathlib.Data.Rat.Sqrt import Mathlib.Data.Real.Sqrt -import Mathlib.RingTheory.Algebraic -import Mathlib.RingTheory.Int.Basic +import Mathlib.RingTheory.Algebraic.Basic import Mathlib.Tactic.IntervalCases /-! diff --git a/Mathlib/FieldTheory/AxGrothendieck.lean b/Mathlib/FieldTheory/AxGrothendieck.lean index b3ce93355bef9..5cfd3ecd0a861 100644 --- a/Mathlib/FieldTheory/AxGrothendieck.lean +++ b/Mathlib/FieldTheory/AxGrothendieck.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import Mathlib.RingTheory.Algebraic +import Mathlib.RingTheory.Algebraic.Basic import Mathlib.Data.Fintype.Card import Mathlib.ModelTheory.Algebra.Field.IsAlgClosed import Mathlib.ModelTheory.Algebra.Ring.Definability diff --git a/Mathlib/FieldTheory/IntermediateField/Algebraic.lean b/Mathlib/FieldTheory/IntermediateField/Algebraic.lean index fadc8889330f0..db5ed7f190e80 100644 --- a/Mathlib/FieldTheory/IntermediateField/Algebraic.lean +++ b/Mathlib/FieldTheory/IntermediateField/Algebraic.lean @@ -4,10 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ import Mathlib.FieldTheory.IntermediateField.Basic -import Mathlib.RingTheory.Algebraic -import Mathlib.FieldTheory.Tower import Mathlib.FieldTheory.Minpoly.Basic +import Mathlib.FieldTheory.Tower import Mathlib.LinearAlgebra.FreeModule.StrongRankCondition +import Mathlib.RingTheory.Algebraic.Integral /-! # Results on finite dimensionality and algebraicity of intermediate fields. diff --git a/Mathlib/FieldTheory/Minpoly/Field.lean b/Mathlib/FieldTheory/Minpoly/Field.lean index 6adb8ecfc42d0..e071240005dbf 100644 --- a/Mathlib/FieldTheory/Minpoly/Field.lean +++ b/Mathlib/FieldTheory/Minpoly/Field.lean @@ -5,7 +5,8 @@ Authors: Riccardo Brasca, Johan Commelin -/ import Mathlib.Algebra.Polynomial.FieldDivision import Mathlib.FieldTheory.Minpoly.Basic -import Mathlib.RingTheory.Algebraic +import Mathlib.RingTheory.Algebraic.Integral +import Mathlib.RingTheory.LocalRing.Basic /-! # Minimal polynomials on an algebra over a field diff --git a/Mathlib/FieldTheory/MvRatFunc/Rank.lean b/Mathlib/FieldTheory/MvRatFunc/Rank.lean index e6f6bcac0babe..8843b5d78d6fc 100644 --- a/Mathlib/FieldTheory/MvRatFunc/Rank.lean +++ b/Mathlib/FieldTheory/MvRatFunc/Rank.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jz Pan -/ import Mathlib.Algebra.MvPolynomial.Cardinal -import Mathlib.RingTheory.Algebraic +import Mathlib.RingTheory.Algebraic.LinearIndependent +import Mathlib.RingTheory.Algebraic.MvPolynomial import Mathlib.RingTheory.Localization.Cardinality import Mathlib.RingTheory.MvPolynomial diff --git a/Mathlib/FieldTheory/SplittingField/Construction.lean b/Mathlib/FieldTheory/SplittingField/Construction.lean index b9eba35bfa239..9177d7e31ecfa 100644 --- a/Mathlib/FieldTheory/SplittingField/Construction.lean +++ b/Mathlib/FieldTheory/SplittingField/Construction.lean @@ -5,6 +5,7 @@ Authors: Chris Hughes -/ import Mathlib.Algebra.CharP.Algebra import Mathlib.FieldTheory.SplittingField.IsSplittingField +import Mathlib.RingTheory.Algebraic.Basic /-! # Splitting fields diff --git a/Mathlib/NumberTheory/DiophantineApproximation.lean b/Mathlib/NumberTheory/DiophantineApproximation.lean index cdccd001950cb..a409528822355 100644 --- a/Mathlib/NumberTheory/DiophantineApproximation.lean +++ b/Mathlib/NumberTheory/DiophantineApproximation.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.ContinuedFractions.Computation.ApproximationCorollaries import Mathlib.Algebra.ContinuedFractions.Computation.Translations import Mathlib.Data.Real.Irrational import Mathlib.RingTheory.Coprime.Lemmas +import Mathlib.RingTheory.Int.Basic import Mathlib.Tactic.Basic /-! diff --git a/Mathlib/RingTheory/Algebraic.lean b/Mathlib/RingTheory/Algebraic/Basic.lean similarity index 61% rename from Mathlib/RingTheory/Algebraic.lean rename to Mathlib/RingTheory/Algebraic/Basic.lean index deafa7a49ce6a..2894940d4858f 100644 --- a/Mathlib/RingTheory/Algebraic.lean +++ b/Mathlib/RingTheory/Algebraic/Basic.lean @@ -3,10 +3,10 @@ Copyright (c) 2019 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ -import Mathlib.RingTheory.IntegralClosure.IsIntegralClosure.Basic -import Mathlib.RingTheory.Polynomial.IntegralNormalization -import Mathlib.RingTheory.LocalRing.Basic -import Mathlib.Algebra.MvPolynomial.Supported +import Mathlib.Algebra.Polynomial.Expand +import Mathlib.Algebra.Polynomial.Roots +import Mathlib.RingTheory.Algebraic.Defs +import Mathlib.RingTheory.Polynomial.Tower /-! # Algebraic elements and algebraic extensions @@ -17,39 +17,20 @@ The main result in this file proves transitivity of algebraicity: a tower of algebraic field extensions is algebraic. -/ - universe u v w -open scoped Classical open Polynomial section variable (R : Type u) {A : Type v} [CommRing R] [Ring A] [Algebra R A] -/-- An element of an R-algebra is algebraic over R if it is a root of a nonzero polynomial -with coefficients in R. -/ -@[stacks 09GC "Algebraic elements"] -def IsAlgebraic (x : A) : Prop := - ∃ p : R[X], p ≠ 0 ∧ aeval x p = 0 - -/-- An element of an R-algebra is transcendental over R if it is not algebraic over R. -/ -def Transcendental (x : A) : Prop := - ¬IsAlgebraic R x - @[nontriviality] theorem is_transcendental_of_subsingleton [Subsingleton R] (x : A) : Transcendental R x := fun ⟨p, h, _⟩ => h <| Subsingleton.elim p 0 variable {R} -/-- An element `x` is transcendental over `R` if and only if for any polynomial `p`, -`Polynomial.aeval x p = 0` implies `p = 0`. This is similar to `algebraicIndependent_iff`. -/ -theorem transcendental_iff {x : A} : - Transcendental R x ↔ ∀ p : R[X], aeval x p = 0 → p = 0 := by - rw [Transcendental, IsAlgebraic, not_exists] - congr! 1; tauto - variable (R) in theorem Polynomial.transcendental_X : Transcendental R (X (R := R)) := by simp [transcendental_iff] @@ -96,76 +77,6 @@ theorem Polynomial.transcendental (f : R[X]) (hf : f.natDegree ≠ 0) Transcendental R f := by simpa using (transcendental_X R).aeval f hf hf' -/-- If `E / F` is a field extension, `x` is an element of `E` transcendental over `F`, -then `{(x - a)⁻¹ | a : F}` is linearly independent over `F`. -/ -theorem Transcendental.linearIndependent_sub_inv - {F E : Type*} [Field F] [Field E] [Algebra F E] {x : E} (H : Transcendental F x) : - LinearIndependent F fun a ↦ (x - algebraMap F E a)⁻¹ := by - rw [transcendental_iff] at H - refine linearIndependent_iff'.2 fun s m hm i hi ↦ ?_ - have hnz (a : F) : x - algebraMap F E a ≠ 0 := fun h ↦ - X_sub_C_ne_zero a <| H (.X - .C a) (by simp [h]) - let b := s.prod fun j ↦ x - algebraMap F E j - have h1 : ∀ i ∈ s, m i • (b * (x - algebraMap F E i)⁻¹) = - m i • (s.erase i).prod fun j ↦ x - algebraMap F E j := fun i hi ↦ by - simp_rw [b, ← s.prod_erase_mul _ hi, mul_inv_cancel_right₀ (hnz i)] - replace hm := congr(b * $(hm)) - simp_rw [mul_zero, Finset.mul_sum, mul_smul_comm, Finset.sum_congr rfl h1] at hm - let p : Polynomial F := s.sum fun i ↦ .C (m i) * (s.erase i).prod fun j ↦ .X - .C j - replace hm := congr(Polynomial.aeval i $(H p (by simp_rw [← hm, p, map_sum, map_mul, map_prod, - map_sub, aeval_X, aeval_C, Algebra.smul_def]))) - have h2 : ∀ j ∈ s.erase i, m j * ((s.erase j).prod fun x ↦ i - x) = 0 := fun j hj ↦ by - have := Finset.mem_erase_of_ne_of_mem (Finset.ne_of_mem_erase hj).symm hi - simp_rw [← (s.erase j).prod_erase_mul _ this, sub_self, mul_zero] - simp_rw [map_zero, p, map_sum, map_mul, map_prod, map_sub, aeval_X, - aeval_C, Algebra.id.map_eq_self, ← s.sum_erase_add _ hi, - Finset.sum_eq_zero h2, zero_add] at hm - exact eq_zero_of_ne_zero_of_mul_right_eq_zero (Finset.prod_ne_zero_iff.2 fun j hj ↦ - sub_ne_zero.2 (Finset.ne_of_mem_erase hj).symm) hm - -/-- A subalgebra is algebraic if all its elements are algebraic. -/ -nonrec -def Subalgebra.IsAlgebraic (S : Subalgebra R A) : Prop := - ∀ x ∈ S, IsAlgebraic R x - -variable (R A) - -/-- An algebra is algebraic if all its elements are algebraic. -/ -@[stacks 09GC "Algebraic extensions"] -protected class Algebra.IsAlgebraic : Prop where - isAlgebraic : ∀ x : A, IsAlgebraic R x - -/-- An algebra is transcendental if some element is transcendental. -/ -protected class Algebra.Transcendental : Prop where - transcendental : ∃ x : A, Transcendental R x - -variable {R A} - -lemma Algebra.isAlgebraic_def : Algebra.IsAlgebraic R A ↔ ∀ x : A, IsAlgebraic R x := - ⟨fun ⟨h⟩ ↦ h, fun h ↦ ⟨h⟩⟩ - -lemma Algebra.transcendental_def : Algebra.Transcendental R A ↔ ∃ x : A, Transcendental R x := - ⟨fun ⟨h⟩ ↦ h, fun h ↦ ⟨h⟩⟩ - -theorem Algebra.transcendental_iff_not_isAlgebraic : - Algebra.Transcendental R A ↔ ¬ Algebra.IsAlgebraic R A := by - simp [isAlgebraic_def, transcendental_def, Transcendental] - -/-- A subalgebra is algebraic if and only if it is algebraic as an algebra. -/ -theorem Subalgebra.isAlgebraic_iff (S : Subalgebra R A) : - S.IsAlgebraic ↔ Algebra.IsAlgebraic R S := by - delta Subalgebra.IsAlgebraic - rw [Subtype.forall', Algebra.isAlgebraic_def] - refine forall_congr' fun x => exists_congr fun p => and_congr Iff.rfl ?_ - have h : Function.Injective S.val := Subtype.val_injective - conv_rhs => rw [← h.eq_iff, map_zero] - rw [← aeval_algHom_apply, S.val_apply] - -/-- An algebra is algebraic if and only if it is algebraic as a subalgebra. -/ -theorem Algebra.isAlgebraic_iff : Algebra.IsAlgebraic R A ↔ (⊤ : Subalgebra R A).IsAlgebraic := by - delta Subalgebra.IsAlgebraic - simp only [Algebra.isAlgebraic_def, Algebra.mem_top, forall_prop_of_true] - theorem isAlgebraic_iff_not_injective {x : A} : IsAlgebraic R x ↔ ¬Function.Injective (Polynomial.aeval x : R[X] →ₐ[R] A) := by simp only [IsAlgebraic, injective_iff_map_eq_zero, not_forall, and_comm, exists_prop] @@ -190,13 +101,6 @@ variable {R : Type u} {S : Type*} {A : Type v} [CommRing R] variable [CommRing S] [Ring A] [Algebra R A] [Algebra R S] [Algebra S A] variable [IsScalarTower R S A] -/-- An integral element of an algebra is algebraic. -/ -theorem IsIntegral.isAlgebraic [Nontrivial R] {x : A} : IsIntegral R x → IsAlgebraic R x := - fun ⟨p, hp, hpx⟩ => ⟨p, hp.ne_zero, hpx⟩ - -instance Algebra.IsIntegral.isAlgebraic [Nontrivial R] [Algebra.IsIntegral R A] : - Algebra.IsAlgebraic R A := ⟨fun a ↦ (Algebra.IsIntegral.isIntegral a).isAlgebraic⟩ - theorem isAlgebraic_zero [Nontrivial R] : IsAlgebraic R (0 : A) := ⟨_, X_ne_zero, aeval_X 0⟩ @@ -443,49 +347,6 @@ alias ⟨_, IsAlgebraic.inv⟩ := IsAlgebraic.inv_iff end zero_ne_one -section Field - -variable {K : Type u} {A : Type v} [Field K] [Ring A] [Algebra K A] - -/-- An element of an algebra over a field is algebraic if and only if it is integral. -/ -theorem isAlgebraic_iff_isIntegral {x : A} : IsAlgebraic K x ↔ IsIntegral K x := by - refine ⟨?_, IsIntegral.isAlgebraic⟩ - rintro ⟨p, hp, hpx⟩ - refine ⟨_, monic_mul_leadingCoeff_inv hp, ?_⟩ - rw [← aeval_def, map_mul, hpx, zero_mul] - -protected theorem Algebra.isAlgebraic_iff_isIntegral : - Algebra.IsAlgebraic K A ↔ Algebra.IsIntegral K A := by - rw [Algebra.isAlgebraic_def, Algebra.isIntegral_def, - forall_congr' fun _ ↦ isAlgebraic_iff_isIntegral] - -alias ⟨IsAlgebraic.isIntegral, _⟩ := isAlgebraic_iff_isIntegral - -/-- This used to be an `alias` of `Algebra.isAlgebraic_iff_isIntegral` but that would make -`Algebra.IsAlgebraic K A` an explicit parameter instead of instance implicit. -/ -protected instance Algebra.IsAlgebraic.isIntegral [Algebra.IsAlgebraic K A] : - Algebra.IsIntegral K A := Algebra.isAlgebraic_iff_isIntegral.mp ‹_› - -variable (K) in -theorem Algebra.IsAlgebraic.of_isIntegralClosure (B C : Type*) - [CommRing B] [CommRing C] [Algebra K B] [Algebra K C] [Algebra B C] - [IsScalarTower K B C] [IsIntegralClosure B K C] : Algebra.IsAlgebraic K B := - Algebra.isAlgebraic_iff_isIntegral.mpr (IsIntegralClosure.isIntegral_algebra K C) - -/-- If `K` is a field, `r : A` and `f : K[X]`, then `Polynomial.aeval r f` is -transcendental over `K` if and only if `r` and `f` are both transcendental over `K`. -See also `Transcendental.aeval_of_transcendental` and `Transcendental.of_aeval`. -/ -@[simp] -theorem transcendental_aeval_iff {r : A} {f : K[X]} : - Transcendental K (Polynomial.aeval r f) ↔ Transcendental K r ∧ Transcendental K f := by - refine ⟨fun h ↦ ⟨?_, h.of_aeval⟩, fun ⟨h1, h2⟩ ↦ h1.aeval_of_transcendental h2⟩ - rw [Transcendental] at h ⊢ - contrapose! h - rw [isAlgebraic_iff_isIntegral] at h ⊢ - exact .of_mem_of_fg _ h.fg_adjoin_singleton _ (aeval_mem_adjoin_singleton _ _) - -end Field - section variable {K L R S A : Type*} @@ -572,17 +433,7 @@ theorem Transcendental.of_tower_top {x : A} (h : Transcendental L x) : theorem Algebra.IsAlgebraic.tower_top [Algebra.IsAlgebraic K A] : Algebra.IsAlgebraic L A := Algebra.IsAlgebraic.extendScalars (algebraMap K L).injective -variable (K) - -theorem IsAlgebraic.of_finite (e : A) [FiniteDimensional K A] : IsAlgebraic K e := - (IsIntegral.of_finite K e).isAlgebraic - -variable (A) - -/-- A field extension is algebraic if it is finite. -/ -@[stacks 09GG "first part"] -instance Algebra.IsAlgebraic.of_finite [FiniteDimensional K A] : Algebra.IsAlgebraic K A := - (IsIntegral.of_finite K A).isAlgebraic +variable (K) (A) theorem Algebra.IsAlgebraic.tower_bot (K L A : Type*) [CommRing K] [Field L] [Ring A] [Algebra K L] [Algebra L A] [Algebra K A] [IsScalarTower K L A] @@ -594,22 +445,6 @@ end Field end Ring -section CommRing - -variable [Field K] [Field L] [Ring A] -variable [Algebra K L] [Algebra L A] [Algebra K A] [IsScalarTower K L A] - -/-- If L is an algebraic field extension of K and A is an algebraic algebra over L, -then A is algebraic over K. -/ -@[stacks 09GJ] -protected theorem Algebra.IsAlgebraic.trans - [L_alg : Algebra.IsAlgebraic K L] [A_alg : Algebra.IsAlgebraic L A] : - Algebra.IsAlgebraic K A := by - rw [Algebra.isAlgebraic_iff_isIntegral] at L_alg A_alg ⊢ - exact Algebra.IsIntegral.trans L - -end CommRing - section NoZeroSMulDivisors namespace Algebra.IsAlgebraic @@ -659,23 +494,6 @@ end Algebra.IsAlgebraic end NoZeroSMulDivisors -section Field - -variable [Field K] [Field L] -variable [Algebra K L] - -theorem AlgHom.bijective [FiniteDimensional K L] (ϕ : L →ₐ[K] L) : Function.Bijective ϕ := - (Algebra.IsAlgebraic.of_finite K L).algHom_bijective ϕ - -variable (K L) - -/-- Bijection between algebra equivalences and algebra homomorphisms -/ -noncomputable abbrev algEquivEquivAlgHom [FiniteDimensional K L] : - (L ≃ₐ[K] L) ≃* (L →ₐ[K] L) := - Algebra.IsAlgebraic.algEquivEquivAlgHom K L - -end Field - end section @@ -720,17 +538,6 @@ end variable {R S : Type*} [CommRing R] [IsDomain R] [CommRing S] -theorem exists_integral_multiple [Algebra R S] {z : S} (hz : IsAlgebraic R z) - (inj : ∀ x, algebraMap R S x = 0 → x = 0) : - ∃ᵉ (x : integralClosure R S) (y ≠ (0 : R)), z * algebraMap R S y = x := by - rcases hz with ⟨p, p_ne_zero, px⟩ - set a := p.leadingCoeff - have a_ne_zero : a ≠ 0 := mt Polynomial.leadingCoeff_eq_zero.mp p_ne_zero - have x_integral : IsIntegral R (z * algebraMap R S a) := - ⟨p.integralNormalization, monic_integralNormalization p_ne_zero, - integralNormalization_aeval_eq_zero px inj⟩ - exact ⟨⟨_, x_integral⟩, a, a_ne_zero, rfl⟩ - section Field variable {K L : Type*} [Field K] [Field L] [Algebra K L] (A : Subalgebra K L) @@ -792,166 +599,6 @@ theorem Subalgebra.isField_of_algebraic [Algebra.IsAlgebraic K L] : IsField A := end Field -section Pi - -variable (R' : Type u) (S' : Type v) (T' : Type w) - -/-- This is not an instance as it forms a diamond with `Pi.instSMul`. - -See the `instance_diamonds` test for details. -/ -def Polynomial.hasSMulPi [Semiring R'] [SMul R' S'] : SMul R'[X] (R' → S') := - ⟨fun p f x => eval x p • f x⟩ - -/-- This is not an instance as it forms a diamond with `Pi.instSMul`. - -See the `instance_diamonds` test for details. -/ -noncomputable def Polynomial.hasSMulPi' [CommSemiring R'] [Semiring S'] [Algebra R' S'] - [SMul S' T'] : SMul R'[X] (S' → T') := - ⟨fun p f x => aeval x p • f x⟩ - -attribute [local instance] Polynomial.hasSMulPi Polynomial.hasSMulPi' - -@[simp] -theorem polynomial_smul_apply [Semiring R'] [SMul R' S'] (p : R'[X]) (f : R' → S') (x : R') : - (p • f) x = eval x p • f x := - rfl - -@[simp] -theorem polynomial_smul_apply' [CommSemiring R'] [Semiring S'] [Algebra R' S'] [SMul S' T'] - (p : R'[X]) (f : S' → T') (x : S') : (p • f) x = aeval x p • f x := - rfl - -variable [CommSemiring R'] [CommSemiring S'] [CommSemiring T'] [Algebra R' S'] [Algebra S' T'] - --- Porting note: the proofs in this definition used `funext` in term-mode, but I was not able --- to get them to work anymore. -/-- This is not an instance for the same reasons as `Polynomial.hasSMulPi'`. -/ -noncomputable def Polynomial.algebraPi : Algebra R'[X] (S' → T') := - { Polynomial.hasSMulPi' R' S' T' with - toFun := fun p z => algebraMap S' T' (aeval z p) - map_one' := by - funext z - simp only [Polynomial.aeval_one, Pi.one_apply, map_one] - map_mul' := fun f g => by - funext z - simp only [Pi.mul_apply, map_mul] - map_zero' := by - funext z - simp only [Polynomial.aeval_zero, Pi.zero_apply, map_zero] - map_add' := fun f g => by - funext z - simp only [Polynomial.aeval_add, Pi.add_apply, map_add] - commutes' := fun p f => by - funext z - exact mul_comm _ _ - smul_def' := fun p f => by - funext z - simp only [polynomial_smul_apply', Algebra.algebraMap_eq_smul_one, RingHom.coe_mk, - MonoidHom.coe_mk, OneHom.coe_mk, Pi.mul_apply, Algebra.smul_mul_assoc, one_mul] } - -attribute [local instance] Polynomial.algebraPi - -@[simp] -theorem Polynomial.algebraMap_pi_eq_aeval : - (algebraMap R'[X] (S' → T') : R'[X] → S' → T') = fun p z => algebraMap _ _ (aeval z p) := - rfl - -@[simp] -theorem Polynomial.algebraMap_pi_self_eq_eval : - (algebraMap R'[X] (R' → R') : R'[X] → R' → R') = fun p z => eval z p := - rfl - -end Pi - -namespace MvPolynomial - -variable {σ : Type*} (R : Type*) [CommRing R] - --- TODO: move to suitable place -private theorem rename_polynomial_aeval_X - {σ τ R : Type*} [CommSemiring R] (f : σ → τ) (i : σ) (p : R[X]) : - rename f (Polynomial.aeval (X i) p) = Polynomial.aeval (X (f i) : MvPolynomial τ R) p := by - rw [← AlgHom.comp_apply] - congr 1; ext1; simp - -theorem transcendental_supported_polynomial_aeval_X {i : σ} {s : Set σ} (h : i ∉ s) - {f : R[X]} (hf : Transcendental R f) : - Transcendental (supported R s) (Polynomial.aeval (X i : MvPolynomial σ R) f) := by - rw [transcendental_iff_injective] at hf ⊢ - let g := MvPolynomial.mapAlgHom (R := R) (σ := s) (Polynomial.aeval (R := R) f) - replace hf : Function.Injective g := MvPolynomial.map_injective _ hf - let u := (Subalgebra.val _).comp - ((optionEquivRight R s).symm |>.trans - (renameEquiv R (Set.subtypeInsertEquivOption h).symm) |>.trans - (supportedEquivMvPolynomial _).symm).toAlgHom |>.comp - g |>.comp - ((optionEquivLeft R s).symm.trans (optionEquivRight R s)).toAlgHom - let v := ((Polynomial.aeval (R := supported R s) - (Polynomial.aeval (X i : MvPolynomial σ R) f)).restrictScalars R).comp - (Polynomial.mapAlgEquiv (supportedEquivMvPolynomial s).symm).toAlgHom - replace hf : Function.Injective u := by - simp only [AlgEquiv.toAlgHom_eq_coe, AlgHom.coe_comp, Subalgebra.coe_val, - AlgHom.coe_coe, AlgEquiv.coe_trans, Function.comp_assoc, u] - apply Subtype.val_injective.comp - simp only [EquivLike.comp_injective] - apply hf.comp - simp only [EquivLike.comp_injective, EquivLike.injective] - have h1 : Polynomial.aeval (X i : MvPolynomial σ R) = ((Subalgebra.val _).comp - (supportedEquivMvPolynomial _).symm.toAlgHom |>.comp - (Polynomial.aeval (X ⟨i, s.mem_insert i⟩ : MvPolynomial ↑(insert i s) R))) := by - ext1; simp - have h2 : u = v := by - simp only [u, v, g] - ext1 - · ext1 - simp [Set.subtypeInsertEquivOption, Subalgebra.algebraMap_eq] - · simp [Set.subtypeInsertEquivOption, rename_polynomial_aeval_X, h1] - simpa only [h2, v, AlgEquiv.toAlgHom_eq_coe, AlgHom.coe_comp, AlgHom.coe_coe, - EquivLike.injective_comp, AlgHom.coe_restrictScalars'] using hf - -theorem transcendental_polynomial_aeval_X (i : σ) {f : R[X]} (hf : Transcendental R f) : - Transcendental R (Polynomial.aeval (X i : MvPolynomial σ R) f) := by - have := transcendental_supported_polynomial_aeval_X R (Set.not_mem_empty i) hf - let g := (Algebra.botEquivOfInjective (MvPolynomial.C_injective σ R)).symm.trans - (Subalgebra.equivOfEq _ _ supported_empty).symm - rwa [Transcendental, ← isAlgebraic_ringHom_iff_of_comp_eq g (RingHom.id (MvPolynomial σ R)) - Function.injective_id (by ext1; rfl), RingHom.id_apply, ← Transcendental] - -theorem transcendental_polynomial_aeval_X_iff (i : σ) {f : R[X]} : - Transcendental R (Polynomial.aeval (X i : MvPolynomial σ R) f) ↔ Transcendental R f := by - refine ⟨?_, transcendental_polynomial_aeval_X R i⟩ - simp_rw [Transcendental, not_imp_not] - exact fun h ↦ h.algHom _ - -theorem transcendental_supported_polynomial_aeval_X_iff - [Nontrivial R] {i : σ} {s : Set σ} {f : R[X]} : - Transcendental (supported R s) (Polynomial.aeval (X i : MvPolynomial σ R) f) ↔ - i ∉ s ∧ Transcendental R f := by - refine ⟨fun h ↦ ⟨?_, ?_⟩, fun ⟨h, hf⟩ ↦ transcendental_supported_polynomial_aeval_X R h hf⟩ - · rw [Transcendental] at h - contrapose! h - refine isAlgebraic_algebraMap (⟨Polynomial.aeval (X i) f, ?_⟩ : supported R s) - exact Algebra.adjoin_mono (Set.singleton_subset_iff.2 (Set.mem_image_of_mem _ h)) - (Polynomial.aeval_mem_adjoin_singleton _ _) - · rw [← transcendental_polynomial_aeval_X_iff R i] - refine h.restrictScalars fun _ _ heq ↦ MvPolynomial.C_injective σ R ?_ - simp_rw [← MvPolynomial.algebraMap_eq] - exact congr($(heq).1) - -theorem transcendental_supported_X {i : σ} {s : Set σ} (h : i ∉ s) : - Transcendental (supported R s) (X i : MvPolynomial σ R) := by - simpa using transcendental_supported_polynomial_aeval_X R h (Polynomial.transcendental_X R) - -theorem transcendental_X (i : σ) : Transcendental R (X i : MvPolynomial σ R) := by - simpa using transcendental_polynomial_aeval_X R i (Polynomial.transcendental_X R) - -theorem transcendental_supported_X_iff [Nontrivial R] {i : σ} {s : Set σ} : - Transcendental (supported R s) (X i : MvPolynomial σ R) ↔ i ∉ s := by - simpa [Polynomial.transcendental_X] using - transcendental_supported_polynomial_aeval_X_iff R (i := i) (s := s) (f := Polynomial.X) - -end MvPolynomial - section Infinite theorem Transcendental.infinite {R A : Type*} [CommRing R] [Ring A] [Algebra R A] diff --git a/Mathlib/RingTheory/Algebraic/Cardinality.lean b/Mathlib/RingTheory/Algebraic/Cardinality.lean index 43311bc18a58e..84f3f265104fc 100644 --- a/Mathlib/RingTheory/Algebraic/Cardinality.lean +++ b/Mathlib/RingTheory/Algebraic/Cardinality.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ import Mathlib.Algebra.Polynomial.Cardinal -import Mathlib.RingTheory.Algebraic +import Mathlib.Algebra.Polynomial.Roots +import Mathlib.RingTheory.Algebraic.Defs /-! # Cardinality of algebraic extensions diff --git a/Mathlib/RingTheory/Algebraic/Defs.lean b/Mathlib/RingTheory/Algebraic/Defs.lean new file mode 100644 index 0000000000000..90f80b749d130 --- /dev/null +++ b/Mathlib/RingTheory/Algebraic/Defs.lean @@ -0,0 +1,103 @@ +/- +Copyright (c) 2019 Johan Commelin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin +-/ +import Mathlib.Algebra.Polynomial.AlgebraMap + +/-! +# Algebraic elements and algebraic extensions + +An element of an R-algebra is algebraic over R if it is the root of a nonzero polynomial. +An R-algebra is algebraic over R if and only if all its elements are algebraic over R. + +## Main definitions + +* `IsAlgebraic`: algebraic elements of an algebra. +* `Transcendental`: transcendental elements of an algebra are those that are not algebraic. +* `Subalgebra.IsAlgebraic`: a subalgebra is algebraic if all its elements are algebraic. +* `Algebra.IsAlgebraic`: an algebra is algebraic if all its elements are algebraic. +* `Algebra.Transcendental`: an algebra is transcendental if some element is transcendental. + +## Main results + +* `transcendental_iff`: an element `x : A` is transcendental over `R` iff out of `R[X]` + only the zero polynomial evaluates to 0 at `x`. +* `Subalgebra.isAlgebraic_iff`: a subalgebra is algebraic iff it is algebraic as an algebra. +-/ + +assert_not_exists IsIntegralClosure +assert_not_exists LinearIndependent +assert_not_exists LocalRing +assert_not_exists MvPolynomial + +universe u v w +open Polynomial + +section + +variable (R : Type u) {A : Type v} [CommRing R] [Ring A] [Algebra R A] + +/-- An element of an R-algebra is algebraic over R if it is a root of a nonzero polynomial +with coefficients in R. -/ +@[stacks 09GC "Algebraic elements"] +def IsAlgebraic (x : A) : Prop := + ∃ p : R[X], p ≠ 0 ∧ aeval x p = 0 + +/-- An element of an R-algebra is transcendental over R if it is not algebraic over R. -/ +def Transcendental (x : A) : Prop := + ¬IsAlgebraic R x + +variable {R} + +/-- An element `x` is transcendental over `R` if and only if for any polynomial `p`, +`Polynomial.aeval x p = 0` implies `p = 0`. This is similar to `algebraicIndependent_iff`. -/ +theorem transcendental_iff {x : A} : + Transcendental R x ↔ ∀ p : R[X], aeval x p = 0 → p = 0 := by + rw [Transcendental, IsAlgebraic, not_exists] + congr! 1; tauto + +/-- A subalgebra is algebraic if all its elements are algebraic. -/ +nonrec +def Subalgebra.IsAlgebraic (S : Subalgebra R A) : Prop := + ∀ x ∈ S, IsAlgebraic R x + +variable (R A) + +/-- An algebra is algebraic if all its elements are algebraic. -/ +@[stacks 09GC "Algebraic extensions"] +protected class Algebra.IsAlgebraic : Prop where + isAlgebraic : ∀ x : A, IsAlgebraic R x + +/-- An algebra is transcendental if some element is transcendental. -/ +protected class Algebra.Transcendental : Prop where + transcendental : ∃ x : A, Transcendental R x + +variable {R A} + +lemma Algebra.isAlgebraic_def : Algebra.IsAlgebraic R A ↔ ∀ x : A, IsAlgebraic R x := + ⟨fun ⟨h⟩ ↦ h, fun h ↦ ⟨h⟩⟩ + +lemma Algebra.transcendental_def : Algebra.Transcendental R A ↔ ∃ x : A, Transcendental R x := + ⟨fun ⟨h⟩ ↦ h, fun h ↦ ⟨h⟩⟩ + +theorem Algebra.transcendental_iff_not_isAlgebraic : + Algebra.Transcendental R A ↔ ¬ Algebra.IsAlgebraic R A := by + simp [isAlgebraic_def, transcendental_def, Transcendental] + +/-- A subalgebra is algebraic if and only if it is algebraic as an algebra. -/ +theorem Subalgebra.isAlgebraic_iff (S : Subalgebra R A) : + S.IsAlgebraic ↔ Algebra.IsAlgebraic R S := by + delta Subalgebra.IsAlgebraic + rw [Subtype.forall', Algebra.isAlgebraic_def] + refine forall_congr' fun x => exists_congr fun p => and_congr Iff.rfl ?_ + have h : Function.Injective S.val := Subtype.val_injective + conv_rhs => rw [← h.eq_iff, map_zero] + rw [← aeval_algHom_apply, S.val_apply] + +/-- An algebra is algebraic if and only if it is algebraic as a subalgebra. -/ +theorem Algebra.isAlgebraic_iff : Algebra.IsAlgebraic R A ↔ (⊤ : Subalgebra R A).IsAlgebraic := by + delta Subalgebra.IsAlgebraic + simp only [Algebra.isAlgebraic_def, Algebra.mem_top, forall_prop_of_true] + +end diff --git a/Mathlib/RingTheory/Algebraic/Integral.lean b/Mathlib/RingTheory/Algebraic/Integral.lean new file mode 100644 index 0000000000000..f5683fd67a76b --- /dev/null +++ b/Mathlib/RingTheory/Algebraic/Integral.lean @@ -0,0 +1,161 @@ +/- +Copyright (c) 2019 Johan Commelin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin +-/ +import Mathlib.RingTheory.Algebraic.Basic +import Mathlib.RingTheory.IntegralClosure.IsIntegralClosure.Basic +import Mathlib.RingTheory.Polynomial.IntegralNormalization + +/-! +# Algebraic elements and integral elements + +This file relates algebraic and integral elements of an algebra, by proving every integral element +is algebraic and that every algebraic element over a field is integral. + +## Main results + +* `IsIntegral.isAlgebraic`, `Algebra.IsIntegral.isAlgebraic`: integral implies algebraic. +* `isAlgebraic_iff_isIntegral`, `Algebra.isAlgebraic_iff_isIntegral`: integral iff algebraic + over a field. +* `IsAlgebraic.of_finite`, `Algebra.IsAlgebraic.of_finite`: finite-dimensional (as module) implies + algebraic. +* `exists_integral_multiple`: an algebraic element has a multiple which is integral +-/ + +assert_not_exists LocalRing + +universe u v w + +open Polynomial + +section zero_ne_one + +variable {R : Type u} {S : Type*} {A : Type v} [CommRing R] +variable [CommRing S] [Ring A] [Algebra R A] [Algebra R S] [Algebra S A] +variable [IsScalarTower R S A] + +/-- An integral element of an algebra is algebraic. -/ +theorem IsIntegral.isAlgebraic [Nontrivial R] {x : A} : IsIntegral R x → IsAlgebraic R x := + fun ⟨p, hp, hpx⟩ => ⟨p, hp.ne_zero, hpx⟩ + +instance Algebra.IsIntegral.isAlgebraic [Nontrivial R] [Algebra.IsIntegral R A] : + Algebra.IsAlgebraic R A := ⟨fun a ↦ (Algebra.IsIntegral.isIntegral a).isAlgebraic⟩ + +end zero_ne_one + +section Field + +variable {K : Type u} {A : Type v} [Field K] [Ring A] [Algebra K A] + +/-- An element of an algebra over a field is algebraic if and only if it is integral. -/ +theorem isAlgebraic_iff_isIntegral {x : A} : IsAlgebraic K x ↔ IsIntegral K x := by + refine ⟨?_, IsIntegral.isAlgebraic⟩ + rintro ⟨p, hp, hpx⟩ + refine ⟨_, monic_mul_leadingCoeff_inv hp, ?_⟩ + rw [← aeval_def, map_mul, hpx, zero_mul] + +protected theorem Algebra.isAlgebraic_iff_isIntegral : + Algebra.IsAlgebraic K A ↔ Algebra.IsIntegral K A := by + rw [Algebra.isAlgebraic_def, Algebra.isIntegral_def, + forall_congr' fun _ ↦ isAlgebraic_iff_isIntegral] + +alias ⟨IsAlgebraic.isIntegral, _⟩ := isAlgebraic_iff_isIntegral + +/-- This used to be an `alias` of `Algebra.isAlgebraic_iff_isIntegral` but that would make +`Algebra.IsAlgebraic K A` an explicit parameter instead of instance implicit. -/ +protected instance Algebra.IsAlgebraic.isIntegral [Algebra.IsAlgebraic K A] : + Algebra.IsIntegral K A := Algebra.isAlgebraic_iff_isIntegral.mp ‹_› + +variable (K) in +theorem Algebra.IsAlgebraic.of_isIntegralClosure (B C : Type*) + [CommRing B] [CommRing C] [Algebra K B] [Algebra K C] [Algebra B C] + [IsScalarTower K B C] [IsIntegralClosure B K C] : Algebra.IsAlgebraic K B := + Algebra.isAlgebraic_iff_isIntegral.mpr (IsIntegralClosure.isIntegral_algebra K C) + +end Field + +section + +variable (K L : Type*) {R S A : Type*} + +section Ring + +section Field + +variable [Field K] [Field L] [Ring A] +variable [Algebra K L] [Algebra L A] [Algebra K A] [IsScalarTower K L A] + +theorem IsAlgebraic.of_finite (e : A) [FiniteDimensional K A] : IsAlgebraic K e := + (IsIntegral.of_finite K e).isAlgebraic + +variable (A) + +/-- A field extension is algebraic if it is finite. -/ +@[stacks 09GG "first part"] +instance Algebra.IsAlgebraic.of_finite [FiniteDimensional K A] : Algebra.IsAlgebraic K A := + (IsIntegral.of_finite K A).isAlgebraic + +end Field + +end Ring + +section CommRing + +variable {K L} [Field K] [Field L] [Ring A] +variable [Algebra K L] [Algebra L A] [Algebra K A] [IsScalarTower K L A] + +/-- If L is an algebraic field extension of K and A is an algebraic algebra over L, +then A is algebraic over K. -/ +@[stacks 09GJ] +protected theorem Algebra.IsAlgebraic.trans + [L_alg : Algebra.IsAlgebraic K L] [A_alg : Algebra.IsAlgebraic L A] : + Algebra.IsAlgebraic K A := by + rw [Algebra.isAlgebraic_iff_isIntegral] at L_alg A_alg ⊢ + exact Algebra.IsIntegral.trans L + +end CommRing + +section Field + +variable {K L} [Field K] [Ring A] [Algebra K A] + +/-- If `K` is a field, `r : A` and `f : K[X]`, then `Polynomial.aeval r f` is +transcendental over `K` if and only if `r` and `f` are both transcendental over `K`. +See also `Transcendental.aeval_of_transcendental` and `Transcendental.of_aeval`. -/ +@[simp] +theorem transcendental_aeval_iff {r : A} {f : K[X]} : + Transcendental K (Polynomial.aeval r f) ↔ Transcendental K r ∧ Transcendental K f := by + refine ⟨fun h ↦ ⟨?_, h.of_aeval⟩, fun ⟨h1, h2⟩ ↦ h1.aeval_of_transcendental h2⟩ + rw [Transcendental] at h ⊢ + contrapose! h + rw [isAlgebraic_iff_isIntegral] at h ⊢ + exact .of_mem_of_fg _ h.fg_adjoin_singleton _ (aeval_mem_adjoin_singleton _ _) + +variable [Field L] [Algebra K L] + +theorem AlgHom.bijective [FiniteDimensional K L] (ϕ : L →ₐ[K] L) : Function.Bijective ϕ := + (Algebra.IsAlgebraic.of_finite K L).algHom_bijective ϕ + +variable (K L) in +/-- Bijection between algebra equivalences and algebra homomorphisms -/ +noncomputable abbrev algEquivEquivAlgHom [FiniteDimensional K L] : + (L ≃ₐ[K] L) ≃* (L →ₐ[K] L) := + Algebra.IsAlgebraic.algEquivEquivAlgHom K L + +end Field + +end + +variable {R S : Type*} [CommRing R] [IsDomain R] [CommRing S] + +theorem exists_integral_multiple [Algebra R S] {z : S} (hz : IsAlgebraic R z) + (inj : ∀ x, algebraMap R S x = 0 → x = 0) : + ∃ᵉ (x : integralClosure R S) (y ≠ (0 : R)), z * algebraMap R S y = x := by + rcases hz with ⟨p, p_ne_zero, px⟩ + set a := p.leadingCoeff + have a_ne_zero : a ≠ 0 := mt Polynomial.leadingCoeff_eq_zero.mp p_ne_zero + have x_integral : IsIntegral R (z * algebraMap R S a) := + ⟨p.integralNormalization, monic_integralNormalization p_ne_zero, + integralNormalization_aeval_eq_zero px inj⟩ + exact ⟨⟨_, x_integral⟩, a, a_ne_zero, rfl⟩ diff --git a/Mathlib/RingTheory/Algebraic/LinearIndependent.lean b/Mathlib/RingTheory/Algebraic/LinearIndependent.lean new file mode 100644 index 0000000000000..4104fb48d25b5 --- /dev/null +++ b/Mathlib/RingTheory/Algebraic/LinearIndependent.lean @@ -0,0 +1,50 @@ +/- +Copyright (c) 2019 Johan Commelin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin +-/ +import Mathlib.LinearAlgebra.LinearIndependent +import Mathlib.RingTheory.Algebraic.Defs + +/-! +# Linear independence of transcendental elements + +## Main result + +* `Transcendental.linearIndependent_sub_inv`: let `x : E` transcendental over `F`, + then `{(x - a)⁻¹ | a : F}` is linearly independent over `F`. +-/ + +open Polynomial + +section + +/-- If `E / F` is a field extension, `x` is an element of `E` transcendental over `F`, +then `{(x - a)⁻¹ | a : F}` is linearly independent over `F`. -/ +theorem Transcendental.linearIndependent_sub_inv + {F E : Type*} [Field F] [Field E] [Algebra F E] {x : E} (H : Transcendental F x) : + LinearIndependent F fun a ↦ (x - algebraMap F E a)⁻¹ := by + classical + rw [transcendental_iff] at H + refine linearIndependent_iff'.2 fun s m hm i hi ↦ ?_ + have hnz (a : F) : x - algebraMap F E a ≠ 0 := fun h ↦ + X_sub_C_ne_zero a <| H (.X - .C a) (by simp [h]) + let b := s.prod fun j ↦ x - algebraMap F E j + have h1 : ∀ i ∈ s, m i • (b * (x - algebraMap F E i)⁻¹) = + m i • (s.erase i).prod fun j ↦ x - algebraMap F E j := fun i hi ↦ by + simp_rw [b, ← s.prod_erase_mul _ hi, mul_inv_cancel_right₀ (hnz i)] + replace hm := congr(b * $(hm)) + simp_rw [mul_zero, Finset.mul_sum, mul_smul_comm, Finset.sum_congr rfl h1] at hm + let p : Polynomial F := s.sum fun i ↦ .C (m i) * (s.erase i).prod fun j ↦ .X - .C j + replace hm := congr(Polynomial.aeval i $(H p (by simp_rw [← hm, p, map_sum, map_mul, map_prod, + map_sub, aeval_X, aeval_C, Algebra.smul_def]))) + have h2 : ∀ j ∈ s.erase i, m j * ((s.erase j).prod fun x ↦ i - x) = 0 := fun j hj ↦ by + have := Finset.mem_erase_of_ne_of_mem (Finset.ne_of_mem_erase hj).symm hi + simp_rw [← (s.erase j).prod_erase_mul _ this, sub_self, mul_zero] + simp_rw [map_zero, p, map_sum, map_mul, map_prod, map_sub, aeval_X, + aeval_C, Algebra.id.map_eq_self, ← s.sum_erase_add _ hi, + Finset.sum_eq_zero h2, zero_add] at hm + exact eq_zero_of_ne_zero_of_mul_right_eq_zero (Finset.prod_ne_zero_iff.2 fun j hj ↦ + sub_ne_zero.2 (Finset.ne_of_mem_erase hj).symm) hm + +end diff --git a/Mathlib/RingTheory/Algebraic/MvPolynomial.lean b/Mathlib/RingTheory/Algebraic/MvPolynomial.lean new file mode 100644 index 0000000000000..649cec44038ad --- /dev/null +++ b/Mathlib/RingTheory/Algebraic/MvPolynomial.lean @@ -0,0 +1,101 @@ +/- +Copyright (c) 2019 Johan Commelin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin +-/ +import Mathlib.Algebra.MvPolynomial.Supported +import Mathlib.RingTheory.Algebraic.Basic + +/-! +# Transcendental elements in `MvPolynomial` + +This file lists some results on some elements in `MvPolynomial σ R` being transcendental +over the base ring `R` and subrings `MvPolynomial.supported` of `MvPolynomial σ R`. +-/ + +universe u v w + +open Polynomial + +namespace MvPolynomial + +variable {σ : Type*} (R : Type*) [CommRing R] + +theorem transcendental_supported_polynomial_aeval_X {i : σ} {s : Set σ} (h : i ∉ s) + {f : R[X]} (hf : Transcendental R f) : + Transcendental (supported R s) (Polynomial.aeval (X i : MvPolynomial σ R) f) := by + classical + rw [transcendental_iff_injective] at hf ⊢ + let g := MvPolynomial.mapAlgHom (R := R) (σ := s) (Polynomial.aeval (R := R) f) + replace hf : Function.Injective g := MvPolynomial.map_injective _ hf + let u := (Subalgebra.val _).comp + ((optionEquivRight R s).symm |>.trans + (renameEquiv R (Set.subtypeInsertEquivOption h).symm) |>.trans + (supportedEquivMvPolynomial _).symm).toAlgHom |>.comp + g |>.comp + ((optionEquivLeft R s).symm.trans (optionEquivRight R s)).toAlgHom + let v := ((Polynomial.aeval (R := supported R s) + (Polynomial.aeval (X i : MvPolynomial σ R) f)).restrictScalars R).comp + (Polynomial.mapAlgEquiv (supportedEquivMvPolynomial s).symm).toAlgHom + replace hf : Function.Injective u := by + simp only [AlgEquiv.toAlgHom_eq_coe, AlgHom.coe_comp, Subalgebra.coe_val, + AlgHom.coe_coe, AlgEquiv.coe_trans, Function.comp_assoc, u] + apply Subtype.val_injective.comp + simp only [EquivLike.comp_injective] + apply hf.comp + simp only [EquivLike.comp_injective, EquivLike.injective] + have h1 : Polynomial.aeval (X i : MvPolynomial σ R) = ((Subalgebra.val _).comp + (supportedEquivMvPolynomial _).symm.toAlgHom |>.comp + (Polynomial.aeval (X ⟨i, s.mem_insert i⟩ : MvPolynomial ↑(insert i s) R))) := by + ext1; simp + have h2 : u = v := by + simp only [u, v, g] + ext1 + · ext1 + simp [Set.subtypeInsertEquivOption, Subalgebra.algebraMap_eq] + · simp [Set.subtypeInsertEquivOption, h1] + simpa only [h2, v, AlgEquiv.toAlgHom_eq_coe, AlgHom.coe_comp, AlgHom.coe_coe, + EquivLike.injective_comp, AlgHom.coe_restrictScalars'] using hf + +theorem transcendental_polynomial_aeval_X (i : σ) {f : R[X]} (hf : Transcendental R f) : + Transcendental R (Polynomial.aeval (X i : MvPolynomial σ R) f) := by + have := transcendental_supported_polynomial_aeval_X R (Set.not_mem_empty i) hf + let g := (Algebra.botEquivOfInjective (MvPolynomial.C_injective σ R)).symm.trans + (Subalgebra.equivOfEq _ _ supported_empty).symm + rwa [Transcendental, ← isAlgebraic_ringHom_iff_of_comp_eq g (RingHom.id (MvPolynomial σ R)) + Function.injective_id (by ext1; rfl), RingHom.id_apply, ← Transcendental] + +theorem transcendental_polynomial_aeval_X_iff (i : σ) {f : R[X]} : + Transcendental R (Polynomial.aeval (X i : MvPolynomial σ R) f) ↔ Transcendental R f := by + refine ⟨?_, transcendental_polynomial_aeval_X R i⟩ + simp_rw [Transcendental, not_imp_not] + exact fun h ↦ h.algHom _ + +theorem transcendental_supported_polynomial_aeval_X_iff + [Nontrivial R] {i : σ} {s : Set σ} {f : R[X]} : + Transcendental (supported R s) (Polynomial.aeval (X i : MvPolynomial σ R) f) ↔ + i ∉ s ∧ Transcendental R f := by + refine ⟨fun h ↦ ⟨?_, ?_⟩, fun ⟨h, hf⟩ ↦ transcendental_supported_polynomial_aeval_X R h hf⟩ + · rw [Transcendental] at h + contrapose! h + refine isAlgebraic_algebraMap (⟨Polynomial.aeval (X i) f, ?_⟩ : supported R s) + exact Algebra.adjoin_mono (Set.singleton_subset_iff.2 (Set.mem_image_of_mem _ h)) + (Polynomial.aeval_mem_adjoin_singleton _ _) + · rw [← transcendental_polynomial_aeval_X_iff R i] + refine h.restrictScalars fun _ _ heq ↦ MvPolynomial.C_injective σ R ?_ + simp_rw [← MvPolynomial.algebraMap_eq] + exact congr($(heq).1) + +theorem transcendental_supported_X {i : σ} {s : Set σ} (h : i ∉ s) : + Transcendental (supported R s) (X i : MvPolynomial σ R) := by + simpa using transcendental_supported_polynomial_aeval_X R h (Polynomial.transcendental_X R) + +theorem transcendental_X (i : σ) : Transcendental R (X i : MvPolynomial σ R) := by + simpa using transcendental_polynomial_aeval_X R i (Polynomial.transcendental_X R) + +theorem transcendental_supported_X_iff [Nontrivial R] {i : σ} {s : Set σ} : + Transcendental (supported R s) (X i : MvPolynomial σ R) ↔ i ∉ s := by + simpa [Polynomial.transcendental_X] using + transcendental_supported_polynomial_aeval_X_iff R (i := i) (s := s) (f := Polynomial.X) + +end MvPolynomial diff --git a/Mathlib/RingTheory/Algebraic/Pi.lean b/Mathlib/RingTheory/Algebraic/Pi.lean new file mode 100644 index 0000000000000..c09c0c2414404 --- /dev/null +++ b/Mathlib/RingTheory/Algebraic/Pi.lean @@ -0,0 +1,90 @@ +/- +Copyright (c) 2019 Johan Commelin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin +-/ +import Mathlib.Algebra.Polynomial.AlgebraMap + +/-! +# Algebraic functions + +This file defines algebraic functions as the image of the `algebraMap R[X] (R → S)`. +-/ + +assert_not_exists IsIntegralClosure +assert_not_exists LinearIndependent +assert_not_exists LocalRing +assert_not_exists MvPolynomial + +open Polynomial + +section Pi + +variable (R S T : Type*) + +/-- This is not an instance as it forms a diamond with `Pi.instSMul`. + +See the `instance_diamonds` test for details. -/ +def Polynomial.hasSMulPi [Semiring R] [SMul R S] : SMul R[X] (R → S) := + ⟨fun p f x => eval x p • f x⟩ + +/-- This is not an instance as it forms a diamond with `Pi.instSMul`. + +See the `instance_diamonds` test for details. -/ +noncomputable def Polynomial.hasSMulPi' [CommSemiring R] [Semiring S] [Algebra R S] + [SMul S T] : SMul R[X] (S → T) := + ⟨fun p f x => aeval x p • f x⟩ + +attribute [local instance] Polynomial.hasSMulPi Polynomial.hasSMulPi' + +@[simp] +theorem polynomial_smul_apply [Semiring R] [SMul R S] (p : R[X]) (f : R → S) (x : R) : + (p • f) x = eval x p • f x := + rfl + +@[simp] +theorem polynomial_smul_apply' [CommSemiring R] [Semiring S] [Algebra R S] [SMul S T] + (p : R[X]) (f : S → T) (x : S) : (p • f) x = aeval x p • f x := + rfl + +variable [CommSemiring R] [CommSemiring S] [CommSemiring T] [Algebra R S] [Algebra S T] + +-- Porting note: the proofs in this definition used `funext` in term-mode, but I was not able +-- to get them to work anymore. +/-- This is not an instance for the same reasons as `Polynomial.hasSMulPi'`. -/ +noncomputable def Polynomial.algebraPi : Algebra R[X] (S → T) := + { Polynomial.hasSMulPi' R S T with + toFun := fun p z => algebraMap S T (aeval z p) + map_one' := by + funext z + simp only [Polynomial.aeval_one, Pi.one_apply, map_one] + map_mul' := fun f g => by + funext z + simp only [Pi.mul_apply, map_mul] + map_zero' := by + funext z + simp only [Polynomial.aeval_zero, Pi.zero_apply, map_zero] + map_add' := fun f g => by + funext z + simp only [Polynomial.aeval_add, Pi.add_apply, map_add] + commutes' := fun p f => by + funext z + exact mul_comm _ _ + smul_def' := fun p f => by + funext z + simp only [polynomial_smul_apply', Algebra.algebraMap_eq_smul_one, RingHom.coe_mk, + MonoidHom.coe_mk, OneHom.coe_mk, Pi.mul_apply, Algebra.smul_mul_assoc, one_mul] } + +attribute [local instance] Polynomial.algebraPi + +@[simp] +theorem Polynomial.algebraMap_pi_eq_aeval : + (algebraMap R[X] (S → T) : R[X] → S → T) = fun p z => algebraMap _ _ (aeval z p) := + rfl + +@[simp] +theorem Polynomial.algebraMap_pi_self_eq_eval : + (algebraMap R[X] (R → R) : R[X] → R → R) = fun p z => eval z p := + rfl + +end Pi diff --git a/Mathlib/RingTheory/AlgebraicIndependent.lean b/Mathlib/RingTheory/AlgebraicIndependent.lean index a82ccfeeb3732..ba9cd9de0655c 100644 --- a/Mathlib/RingTheory/AlgebraicIndependent.lean +++ b/Mathlib/RingTheory/AlgebraicIndependent.lean @@ -9,6 +9,7 @@ import Mathlib.FieldTheory.Adjoin import Mathlib.FieldTheory.MvRatFunc.Rank import Mathlib.RingTheory.Algebraic.Cardinality import Mathlib.RingTheory.MvPolynomial.Basic +import Mathlib.RingTheory.Algebraic.MvPolynomial /-! # Algebraic Independence diff --git a/Mathlib/RingTheory/Localization/Integral.lean b/Mathlib/RingTheory/Localization/Integral.lean index ad327225f2383..91a698e4d9bae 100644 --- a/Mathlib/RingTheory/Localization/Integral.lean +++ b/Mathlib/RingTheory/Localization/Integral.lean @@ -6,7 +6,7 @@ Authors: Kenny Lau, Mario Carneiro, Johan Commelin, Amelia Livingston, Anne Baan import Mathlib.Algebra.GroupWithZero.NonZeroDivisors import Mathlib.Algebra.Polynomial.Lifts import Mathlib.GroupTheory.MonoidLocalization.Basic -import Mathlib.RingTheory.Algebraic +import Mathlib.RingTheory.Algebraic.Integral import Mathlib.RingTheory.IntegralClosure.Algebra.Basic import Mathlib.RingTheory.Localization.FractionRing import Mathlib.RingTheory.Localization.Integer diff --git a/MathlibTest/instance_diamonds.lean b/MathlibTest/instance_diamonds.lean index 5696071a4fb62..06ab59bc5c51d 100644 --- a/MathlibTest/instance_diamonds.lean +++ b/MathlibTest/instance_diamonds.lean @@ -3,13 +3,14 @@ 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.EuclideanDomain.Field import Mathlib.Algebra.GroupWithZero.Action.Prod import Mathlib.Algebra.GroupWithZero.Action.Units import Mathlib.Algebra.Module.Pi import Mathlib.Algebra.Polynomial.Basic import Mathlib.Data.Complex.Module import Mathlib.Data.ZMod.Basic -import Mathlib.RingTheory.Algebraic +import Mathlib.RingTheory.Algebraic.Pi import Mathlib.RingTheory.TensorProduct.Basic /-! # Tests that instances do not form diamonds -/ From ba7312b5aafeb13465bb9283536208f2158e6150 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 30 Nov 2024 09:42:44 +0000 Subject: [PATCH 433/829] chore: move definition of ENat and PNat out of Order.TypeTags (#19485) These had been moved to `Order.TypeTags` to reduce imports elsewhere, but they are sort of off-topic here, so I've just moved them to their own files. --- Mathlib.lean | 2 ++ Mathlib/Data/Countable/Basic.lean | 1 + Mathlib/Data/ENat/Basic.lean | 1 + Mathlib/Data/ENat/Defs.lean | 38 +++++++++++++++++++++++++ Mathlib/Data/PNat/Defs.lean | 2 +- Mathlib/Data/PNat/Notation.lean | 27 ++++++++++++++++++ Mathlib/Order/TypeTags.lean | 47 ------------------------------- 7 files changed, 70 insertions(+), 48 deletions(-) create mode 100644 Mathlib/Data/ENat/Defs.lean create mode 100644 Mathlib/Data/PNat/Notation.lean diff --git a/Mathlib.lean b/Mathlib.lean index cdccddc0c431c..b2c20556f8317 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2389,6 +2389,7 @@ import Mathlib.Data.ENNReal.Operations import Mathlib.Data.ENNReal.Real import Mathlib.Data.ENat.Basic import Mathlib.Data.ENat.BigOperators +import Mathlib.Data.ENat.Defs import Mathlib.Data.ENat.Lattice import Mathlib.Data.Erased import Mathlib.Data.FP.Basic @@ -2767,6 +2768,7 @@ import Mathlib.Data.PNat.Equiv import Mathlib.Data.PNat.Factors import Mathlib.Data.PNat.Find import Mathlib.Data.PNat.Interval +import Mathlib.Data.PNat.Notation import Mathlib.Data.PNat.Prime import Mathlib.Data.PNat.Xgcd import Mathlib.Data.PSigma.Order diff --git a/Mathlib/Data/Countable/Basic.lean b/Mathlib/Data/Countable/Basic.lean index caecc1a3b4c19..e9840f762cea5 100644 --- a/Mathlib/Data/Countable/Basic.lean +++ b/Mathlib/Data/Countable/Basic.lean @@ -5,6 +5,7 @@ Authors: Yury Kudryashov -/ import Mathlib.Data.Countable.Defs import Mathlib.Data.Fin.Tuple.Basic +import Mathlib.Data.ENat.Defs import Mathlib.Logic.Equiv.Nat /-! diff --git a/Mathlib/Data/ENat/Basic.lean b/Mathlib/Data/ENat/Basic.lean index 663a5824cbe31..f2e67d781c703 100644 --- a/Mathlib/Data/ENat/Basic.lean +++ b/Mathlib/Data/ENat/Basic.lean @@ -6,6 +6,7 @@ Authors: Yury Kudryashov import Mathlib.Algebra.CharZero.Lemmas import Mathlib.Algebra.Order.Ring.WithTop import Mathlib.Algebra.Order.Sub.WithTop +import Mathlib.Data.ENat.Defs import Mathlib.Data.Nat.Cast.Order.Basic import Mathlib.Data.Nat.SuccPred import Mathlib.Order.Nat diff --git a/Mathlib/Data/ENat/Defs.lean b/Mathlib/Data/ENat/Defs.lean new file mode 100644 index 0000000000000..479ef16a06353 --- /dev/null +++ b/Mathlib/Data/ENat/Defs.lean @@ -0,0 +1,38 @@ +/- +Copyright (c) 2018 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro, Simon Hudon, Yury Kudryashov +-/ + +import Mathlib.Data.Nat.Notation +import Mathlib.Order.TypeTags + +/-! # Definition and notation for extended natural numbers -/ + +/-- Extended natural numbers `ℕ∞ = WithTop ℕ`. -/ +def ENat : Type := WithTop ℕ deriving Top, Inhabited + +@[inherit_doc] notation "ℕ∞" => ENat + +namespace ENat + +instance instNatCast : NatCast ℕ∞ := ⟨WithTop.some⟩ + +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11445): new definition copied from `WithTop` +/-- Recursor for `ENat` using the preferred forms `⊤` and `↑a`. -/ +@[elab_as_elim, induction_eliminator, cases_eliminator] +def recTopCoe {C : ℕ∞ → Sort*} (top : C ⊤) (coe : ∀ a : ℕ, C a) : ∀ n : ℕ∞, C n + | none => top + | Option.some a => coe a + +@[simp] +theorem recTopCoe_top {C : ℕ∞ → Sort*} (d : C ⊤) (f : ∀ a : ℕ, C a) : + @recTopCoe C d f ⊤ = d := + rfl + +@[simp] +theorem recTopCoe_coe {C : ℕ∞ → Sort*} (d : C ⊤) (f : ∀ a : ℕ, C a) (x : ℕ) : + @recTopCoe C d f ↑x = f x := + rfl + +end ENat diff --git a/Mathlib/Data/PNat/Defs.lean b/Mathlib/Data/PNat/Defs.lean index afc7437cfeddc..f1b5d5d414da3 100644 --- a/Mathlib/Data/PNat/Defs.lean +++ b/Mathlib/Data/PNat/Defs.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Neil Strickland -/ import Mathlib.Data.Nat.Defs +import Mathlib.Data.PNat.Notation import Mathlib.Order.Basic -import Mathlib.Order.TypeTags import Mathlib.Tactic.Coe import Mathlib.Tactic.Lift import Mathlib.Data.Int.Order.Basic diff --git a/Mathlib/Data/PNat/Notation.lean b/Mathlib/Data/PNat/Notation.lean new file mode 100644 index 0000000000000..be4a3da696f59 --- /dev/null +++ b/Mathlib/Data/PNat/Notation.lean @@ -0,0 +1,27 @@ +/- +Copyright (c) 2018 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro, Simon Hudon, Yury Kudryashov +-/ + +import Mathlib.Data.Nat.Notation + +/-! # Definition and notation for positive natural numbers -/ + +/-- `ℕ+` is the type of positive natural numbers. It is defined as a subtype, + and the VM representation of `ℕ+` is the same as `ℕ` because the proof + is not stored. -/ +def PNat := { n : ℕ // 0 < n } deriving DecidableEq + +@[inherit_doc] +notation "ℕ+" => PNat + +/-- The underlying natural number -/ +@[coe] +def PNat.val : ℕ+ → ℕ := Subtype.val + +instance coePNatNat : Coe ℕ+ ℕ := + ⟨PNat.val⟩ + +instance : Repr ℕ+ := + ⟨fun n n' => reprPrec n.1 n'⟩ diff --git a/Mathlib/Order/TypeTags.lean b/Mathlib/Order/TypeTags.lean index 2c8dee9aa90c8..430818ad4338e 100644 --- a/Mathlib/Order/TypeTags.lean +++ b/Mathlib/Order/TypeTags.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Simon Hudon, Yury Kudryashov -/ import Mathlib.Order.Notation -import Mathlib.Data.Nat.Notation /-! # Order-related type synonyms @@ -103,49 +102,3 @@ theorem recTopCoe_coe {C : WithTop α → Sort*} (d : C ⊤) (f : ∀ a : α, C rfl end WithTop - -/-- Extended natural numbers `ℕ∞ = WithTop ℕ`. -/ -def ENat : Type := WithTop ℕ deriving Top, Inhabited - -@[inherit_doc] notation "ℕ∞" => ENat - -namespace ENat - -instance instNatCast : NatCast ℕ∞ := ⟨WithTop.some⟩ - --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11445): new definition copied from `WithTop` -/-- Recursor for `ENat` using the preferred forms `⊤` and `↑a`. -/ -@[elab_as_elim, induction_eliminator, cases_eliminator] -def recTopCoe {C : ℕ∞ → Sort*} (top : C ⊤) (coe : ∀ a : ℕ, C a) : ∀ n : ℕ∞, C n - | none => top - | Option.some a => coe a - -@[simp] -theorem recTopCoe_top {C : ℕ∞ → Sort*} (d : C ⊤) (f : ∀ a : ℕ, C a) : - @recTopCoe C d f ⊤ = d := - rfl - -@[simp] -theorem recTopCoe_coe {C : ℕ∞ → Sort*} (d : C ⊤) (f : ∀ a : ℕ, C a) (x : ℕ) : - @recTopCoe C d f ↑x = f x := - rfl - -end ENat - -/-- `ℕ+` is the type of positive natural numbers. It is defined as a subtype, - and the VM representation of `ℕ+` is the same as `ℕ` because the proof - is not stored. -/ -def PNat := { n : ℕ // 0 < n } deriving DecidableEq - -@[inherit_doc] -notation "ℕ+" => PNat - -/-- The underlying natural number -/ -@[coe] -def PNat.val : ℕ+ → ℕ := Subtype.val - -instance coePNatNat : Coe ℕ+ ℕ := - ⟨PNat.val⟩ - -instance : Repr ℕ+ := - ⟨fun n n' => reprPrec n.1 n'⟩ From 883af569cadb60cf1e7b203baa14dc4a5c964622 Mon Sep 17 00:00:00 2001 From: FR Date: Sat, 30 Nov 2024 18:25:09 +0000 Subject: [PATCH 434/829] chore: remove `normalizeScaleRoots` (#6183) [Zulip](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/duplicate.20definition) Co-authored-by: negiizhao --- Mathlib/Algebra/Polynomial/Basic.lean | 6 + Mathlib/Algebra/Polynomial/Coeff.lean | 1 + Mathlib/RingTheory/Algebraic/Basic.lean | 2 +- Mathlib/RingTheory/Algebraic/Integral.lean | 7 +- .../DedekindDomain/IntegralClosure.lean | 2 +- .../IsIntegralClosure/Basic.lean | 91 +++----- Mathlib/RingTheory/Localization/Integral.lean | 10 +- .../Polynomial/IntegralNormalization.lean | 194 +++++++++++------- 8 files changed, 157 insertions(+), 156 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/Basic.lean b/Mathlib/Algebra/Polynomial/Basic.lean index 72765a3a6e65a..db335123a7319 100644 --- a/Mathlib/Algebra/Polynomial/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Basic.lean @@ -759,6 +759,12 @@ theorem support_monomial' (n) (a : R) : (monomial n a).support ⊆ singleton n : rw [← ofFinsupp_single, support] exact Finsupp.support_single_subset +theorem support_C {a : R} (h : a ≠ 0) : (C a).support = singleton 0 := + support_monomial 0 h + +theorem support_C_subset (a : R) : (C a).support ⊆ singleton 0 := + support_monomial' 0 a + theorem support_C_mul_X {c : R} (h : c ≠ 0) : Polynomial.support (C c * X) = singleton 1 := by rw [C_mul_X_eq_monomial, support_monomial 1 h] diff --git a/Mathlib/Algebra/Polynomial/Coeff.lean b/Mathlib/Algebra/Polynomial/Coeff.lean index 383ce0d93d741..1997d5ce972f0 100644 --- a/Mathlib/Algebra/Polynomial/Coeff.lean +++ b/Mathlib/Algebra/Polynomial/Coeff.lean @@ -101,6 +101,7 @@ lemma coeff_list_sum_map {ι : Type*} (l : List ι) (f : ι → R[X]) (n : ℕ) (l.map f).sum.coeff n = (l.map (fun a => (f a).coeff n)).sum := by simp_rw [coeff_list_sum, List.map_map, Function.comp_def, lcoeff_apply] +@[simp] theorem coeff_sum [Semiring S] (n : ℕ) (f : ℕ → R → S[X]) : coeff (p.sum f) n = p.sum fun a b => coeff (f a b) n := by rcases p with ⟨⟩ diff --git a/Mathlib/RingTheory/Algebraic/Basic.lean b/Mathlib/RingTheory/Algebraic/Basic.lean index 2894940d4858f..98fbcdfdd5f93 100644 --- a/Mathlib/RingTheory/Algebraic/Basic.lean +++ b/Mathlib/RingTheory/Algebraic/Basic.lean @@ -536,7 +536,7 @@ theorem Algebra.IsAlgebraic.exists_smul_eq_mul [IsDomain S] [Algebra.IsAlgebraic end -variable {R S : Type*} [CommRing R] [IsDomain R] [CommRing S] +variable {R S : Type*} [CommRing R] [CommRing S] section Field diff --git a/Mathlib/RingTheory/Algebraic/Integral.lean b/Mathlib/RingTheory/Algebraic/Integral.lean index f5683fd67a76b..b45ecf7023b2f 100644 --- a/Mathlib/RingTheory/Algebraic/Integral.lean +++ b/Mathlib/RingTheory/Algebraic/Integral.lean @@ -5,7 +5,6 @@ Authors: Johan Commelin -/ import Mathlib.RingTheory.Algebraic.Basic import Mathlib.RingTheory.IntegralClosure.IsIntegralClosure.Basic -import Mathlib.RingTheory.Polynomial.IntegralNormalization /-! # Algebraic elements and integral elements @@ -147,15 +146,15 @@ end Field end -variable {R S : Type*} [CommRing R] [IsDomain R] [CommRing S] +variable {R S : Type*} [CommRing R] [CommRing S] theorem exists_integral_multiple [Algebra R S] {z : S} (hz : IsAlgebraic R z) (inj : ∀ x, algebraMap R S x = 0 → x = 0) : - ∃ᵉ (x : integralClosure R S) (y ≠ (0 : R)), z * algebraMap R S y = x := by + ∃ᵉ (x : integralClosure R S) (y ≠ (0 : R)), algebraMap R S y * z = x := by rcases hz with ⟨p, p_ne_zero, px⟩ set a := p.leadingCoeff have a_ne_zero : a ≠ 0 := mt Polynomial.leadingCoeff_eq_zero.mp p_ne_zero - have x_integral : IsIntegral R (z * algebraMap R S a) := + have x_integral : IsIntegral R (algebraMap R S a * z) := ⟨p.integralNormalization, monic_integralNormalization p_ne_zero, integralNormalization_aeval_eq_zero px inj⟩ exact ⟨⟨_, x_integral⟩, a, a_ne_zero, rfl⟩ diff --git a/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean b/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean index f38e19b607e41..b5c1195243656 100644 --- a/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean +++ b/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean @@ -124,7 +124,7 @@ theorem exists_integral_multiples (s : Finset L) : · rcases this with ⟨x', y', hy', hx'⟩ refine ⟨y * y', mul_ne_zero hy hy', fun x'' hx'' => ?_⟩ rcases Finset.mem_insert.mp hx'' with (rfl | hx'') - · rw [mul_smul, Algebra.smul_def, Algebra.smul_def, mul_comm _ x'', hx'] + · rw [mul_smul, Algebra.smul_def, Algebra.smul_def, hx'] exact isIntegral_algebraMap.mul x'.2 · rw [mul_comm, mul_smul, Algebra.smul_def] exact isIntegral_algebraMap.mul (hs _ hx'') diff --git a/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean b/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean index 1cc2210dce70f..17251c27433cb 100644 --- a/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean +++ b/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean @@ -8,6 +8,7 @@ import Mathlib.LinearAlgebra.FiniteDimensional.Defs import Mathlib.RingTheory.IntegralClosure.IsIntegralClosure.Defs import Mathlib.RingTheory.IntegralClosure.Algebra.Basic import Mathlib.RingTheory.FiniteType +import Mathlib.RingTheory.Polynomial.IntegralNormalization import Mathlib.RingTheory.Polynomial.ScaleRoots /-! @@ -262,81 +263,39 @@ section variable (p : R[X]) (x : S) /-- The monic polynomial whose roots are `p.leadingCoeff * x` for roots `x` of `p`. -/ -noncomputable def normalizeScaleRoots (p : R[X]) : R[X] := - ∑ i ∈ p.support, - monomial i (if i = p.natDegree then 1 else p.coeff i * p.leadingCoeff ^ (p.natDegree - 1 - i)) - -theorem normalizeScaleRoots_coeff_mul_leadingCoeff_pow (i : ℕ) (hp : 1 ≤ natDegree p) : - (normalizeScaleRoots p).coeff i * p.leadingCoeff ^ i = - p.coeff i * p.leadingCoeff ^ (p.natDegree - 1) := by - simp only [normalizeScaleRoots, finset_sum_coeff, coeff_monomial, Finset.sum_ite_eq', one_mul, - zero_mul, mem_support_iff, ite_mul, Ne, ite_not] - split_ifs with h₁ h₂ - · simp [h₁] - · rw [h₂, leadingCoeff, ← pow_succ', tsub_add_cancel_of_le hp] - · rw [mul_assoc, ← pow_add, tsub_add_cancel_of_le] - apply Nat.le_sub_one_of_lt - rw [lt_iff_le_and_ne] - exact ⟨le_natDegree_of_ne_zero h₁, h₂⟩ - -theorem leadingCoeff_smul_normalizeScaleRoots (p : R[X]) : - p.leadingCoeff • normalizeScaleRoots p = scaleRoots p p.leadingCoeff := by - ext - simp only [coeff_scaleRoots, normalizeScaleRoots, coeff_monomial, coeff_smul, Finset.smul_sum, - Ne, Finset.sum_ite_eq', finset_sum_coeff, smul_ite, smul_zero, mem_support_iff] - -- Porting note: added the following `simp only` - simp only [tsub_le_iff_right, smul_eq_mul, mul_ite, mul_one, mul_zero, - Finset.sum_ite_eq', mem_support_iff, ne_eq, ite_not] - split_ifs with h₁ h₂ - · simp [*] - · simp [*] - · rw [mul_comm, mul_assoc, ← pow_succ, tsub_right_comm, - tsub_add_cancel_of_le] - rw [Nat.succ_le_iff] - exact tsub_pos_of_lt (lt_of_le_of_ne (le_natDegree_of_ne_zero h₁) h₂) - -theorem normalizeScaleRoots_support : (normalizeScaleRoots p).support ≤ p.support := by - intro x - contrapose - simp only [not_mem_support_iff, normalizeScaleRoots, finset_sum_coeff, coeff_monomial, - Finset.sum_ite_eq', mem_support_iff, Ne, Classical.not_not, ite_eq_right_iff] - intro h₁ h₂ - exact (h₂ h₁).elim - -theorem normalizeScaleRoots_degree : (normalizeScaleRoots p).degree = p.degree := by - apply le_antisymm - · exact Finset.sup_mono (normalizeScaleRoots_support p) - · rw [← degree_scaleRoots, ← leadingCoeff_smul_normalizeScaleRoots] - exact degree_smul_le _ _ - -theorem normalizeScaleRoots_eval₂_leadingCoeff_mul (h : 1 ≤ p.natDegree) (f : R →+* S) (x : S) : - (normalizeScaleRoots p).eval₂ f (f p.leadingCoeff * x) = - f p.leadingCoeff ^ (p.natDegree - 1) * p.eval₂ f x := by - rw [eval₂_eq_sum_range, eval₂_eq_sum_range, Finset.mul_sum] - apply Finset.sum_congr - · rw [natDegree_eq_of_degree_eq (normalizeScaleRoots_degree p)] - intro n _hn - rw [mul_pow, ← mul_assoc, ← f.map_pow, ← f.map_mul, - normalizeScaleRoots_coeff_mul_leadingCoeff_pow _ _ h, f.map_mul, f.map_pow] - ring - -theorem normalizeScaleRoots_monic (h : p ≠ 0) : (normalizeScaleRoots p).Monic := by - delta Monic leadingCoeff - rw [natDegree_eq_of_degree_eq (normalizeScaleRoots_degree p)] - suffices p = 0 → (0 : R) = 1 by simpa [normalizeScaleRoots, coeff_monomial] - exact fun h' => (h h').elim +@[deprecated (since := "2024-11-30")] +alias normalizeScaleRoots := integralNormalization + +@[deprecated (since := "2024-11-30")] +alias normalizeScaleRoots_coeff_mul_leadingCoeff_pow := + integralNormalization_coeff_mul_leadingCoeff_pow + +@[deprecated (since := "2024-11-30")] +alias leadingCoeff_smul_normalizeScaleRoots := leadingCoeff_smul_integralNormalization + +@[deprecated (since := "2024-11-30")] +alias normalizeScaleRoots_support := support_integralNormalization_subset + +@[deprecated (since := "2024-11-30")] +alias normalizeScaleRoots_degree := integralNormalization_degree + +@[deprecated (since := "2024-11-30")] +alias normalizeScaleRoots_eval₂_leadingCoeff_mul := integralNormalization_eval₂_leadingCoeff_mul + +@[deprecated (since := "2024-11-30")] +alias normalizeScaleRoots_monic := monic_integralNormalization /-- Given a `p : R[X]` and a `x : S` such that `p.eval₂ f x = 0`, `f p.leadingCoeff * x` is integral. -/ theorem RingHom.isIntegralElem_leadingCoeff_mul (h : p.eval₂ f x = 0) : f.IsIntegralElem (f p.leadingCoeff * x) := by by_cases h' : 1 ≤ p.natDegree - · use normalizeScaleRoots p + · use integralNormalization p have : p ≠ 0 := fun h'' => by rw [h'', natDegree_zero] at h' exact Nat.not_succ_le_zero 0 h' - use normalizeScaleRoots_monic p this - rw [normalizeScaleRoots_eval₂_leadingCoeff_mul p h' f x, h, mul_zero] + use monic_integralNormalization this + rw [integralNormalization_eval₂_leadingCoeff_mul h' f x, h, mul_zero] · by_cases hp : p.map f = 0 · apply_fun fun q => coeff q p.natDegree at hp rw [coeff_map, coeff_zero, coeff_natDegree] at hp diff --git a/Mathlib/RingTheory/Localization/Integral.lean b/Mathlib/RingTheory/Localization/Integral.lean index 91a698e4d9bae..595fd2c074e5a 100644 --- a/Mathlib/RingTheory/Localization/Integral.lean +++ b/Mathlib/RingTheory/Localization/Integral.lean @@ -279,7 +279,7 @@ theorem IsIntegral.exists_multiple_integral_of_isLocalization [Algebra Rₘ S] [ end IsIntegral -variable {A K : Type*} [CommRing A] [IsDomain A] +variable {A K : Type*} [CommRing A] namespace IsIntegralClosure @@ -306,14 +306,14 @@ theorem isFractionRing_of_algebraic [Algebra.IsAlgebraic A L] hy (inj _ (by rw [IsScalarTower.algebraMap_apply A C L, h, RingHom.map_zero]))⟩, by simp only - rw [algebraMap_mk', ← IsScalarTower.algebraMap_apply A C L, hxy]⟩ + rw [algebraMap_mk', ← IsScalarTower.algebraMap_apply A C L, mul_comm, hxy]⟩ exists_of_eq := fun {x y} h => ⟨1, by simpa using algebraMap_injective C A L h⟩ } variable (K L) /-- If the field `L` is a finite extension of the fraction field of the integral domain `A`, the integral closure `C` of `A` in `L` has fraction field `L`. -/ -theorem isFractionRing_of_finite_extension [Algebra K L] [IsScalarTower A K L] +theorem isFractionRing_of_finite_extension [IsDomain A] [Algebra K L] [IsScalarTower A K L] [FiniteDimensional K L] : IsFractionRing C L := have : Algebra.IsAlgebraic A L := IsFractionRing.comap_isAlgebraic_iff.mpr (inferInstanceAs (Algebra.IsAlgebraic K L)) @@ -341,8 +341,8 @@ variable (K L) /-- If the field `L` is a finite extension of the fraction field of the integral domain `A`, the integral closure of `A` in `L` has fraction field `L`. -/ -theorem isFractionRing_of_finite_extension [Algebra A L] [Algebra K L] [IsScalarTower A K L] - [FiniteDimensional K L] : IsFractionRing (integralClosure A L) L := +theorem isFractionRing_of_finite_extension [IsDomain A] [Algebra A L] [Algebra K L] + [IsScalarTower A K L] [FiniteDimensional K L] : IsFractionRing (integralClosure A L) L := IsIntegralClosure.isFractionRing_of_finite_extension A K L (integralClosure A L) end integralClosure diff --git a/Mathlib/RingTheory/Polynomial/IntegralNormalization.lean b/Mathlib/RingTheory/Polynomial/IntegralNormalization.lean index 75da4b3b8edc2..046c30220049a 100644 --- a/Mathlib/RingTheory/Polynomial/IntegralNormalization.lean +++ b/Mathlib/RingTheory/Polynomial/IntegralNormalization.lean @@ -1,11 +1,10 @@ /- Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Chris Hughes, Johannes Hölzl, Kim Morrison, Jens Wagemaker +Authors: Chris Hughes, Johannes Hölzl, Kim Morrison, Jens Wagemaker, Andrew Yang, Yuyang Zhao -/ -import Mathlib.Algebra.Polynomial.AlgebraMap -import Mathlib.Algebra.Polynomial.Degree.Lemmas import Mathlib.Algebra.Polynomial.Monic +import Mathlib.RingTheory.Polynomial.ScaleRoots /-! # Theory of monic polynomials @@ -28,115 +27,152 @@ section Semiring variable [Semiring R] -/-- If `f : R[X]` is a nonzero polynomial with root `z`, `integralNormalization f` is +/-- If `p : R[X]` is a nonzero polynomial with root `z`, `integralNormalization p` is a monic polynomial with root `leadingCoeff f * z`. Moreover, `integralNormalization 0 = 0`. -/ -noncomputable def integralNormalization (f : R[X]) : R[X] := - ∑ i ∈ f.support, - monomial i (if f.degree = i then 1 else coeff f i * f.leadingCoeff ^ (f.natDegree - 1 - i)) +noncomputable def integralNormalization (p : R[X]) : R[X] := + p.sum fun i a ↦ + monomial i (if p.degree = i then 1 else a * p.leadingCoeff ^ (p.natDegree - 1 - i)) @[simp] theorem integralNormalization_zero : integralNormalization (0 : R[X]) = 0 := by simp [integralNormalization] -theorem integralNormalization_coeff {f : R[X]} {i : ℕ} : - (integralNormalization f).coeff i = - if f.degree = i then 1 else coeff f i * f.leadingCoeff ^ (f.natDegree - 1 - i) := by - have : f.coeff i = 0 → f.degree ≠ i := fun hc hd => coeff_ne_zero_of_eq_degree hd hc - simp +contextual [integralNormalization, coeff_monomial, this, +@[simp] +theorem integralNormalization_C {x : R} (hx : x ≠ 0) : integralNormalization (C x) = 1 := by + simp [integralNormalization, sum_def, support_C hx, degree_C hx] + +variable {p : R[X]} + +theorem integralNormalization_coeff {i : ℕ} : + (integralNormalization p).coeff i = + if p.degree = i then 1 else coeff p i * p.leadingCoeff ^ (p.natDegree - 1 - i) := by + have : p.coeff i = 0 → p.degree ≠ i := fun hc hd => coeff_ne_zero_of_eq_degree hd hc + simp +contextual [sum_def, integralNormalization, coeff_monomial, this, mem_support_iff] -theorem integralNormalization_support {f : R[X]} : - (integralNormalization f).support ⊆ f.support := by +theorem support_integralNormalization_subset : + (integralNormalization p).support ⊆ p.support := by intro - simp +contextual [integralNormalization, coeff_monomial, mem_support_iff] + simp +contextual [sum_def, integralNormalization, coeff_monomial, mem_support_iff] + +@[deprecated (since := "2024-11-30")] +alias integralNormalization_support := support_integralNormalization_subset -theorem integralNormalization_coeff_degree {f : R[X]} {i : ℕ} (hi : f.degree = i) : - (integralNormalization f).coeff i = 1 := by rw [integralNormalization_coeff, if_pos hi] +theorem integralNormalization_coeff_degree {i : ℕ} (hi : p.degree = i) : + (integralNormalization p).coeff i = 1 := by rw [integralNormalization_coeff, if_pos hi] -theorem integralNormalization_coeff_natDegree {f : R[X]} (hf : f ≠ 0) : - (integralNormalization f).coeff (natDegree f) = 1 := - integralNormalization_coeff_degree (degree_eq_natDegree hf) +theorem integralNormalization_coeff_natDegree (hp : p ≠ 0) : + (integralNormalization p).coeff (natDegree p) = 1 := + integralNormalization_coeff_degree (degree_eq_natDegree hp) -theorem integralNormalization_coeff_ne_degree {f : R[X]} {i : ℕ} (hi : f.degree ≠ i) : - coeff (integralNormalization f) i = coeff f i * f.leadingCoeff ^ (f.natDegree - 1 - i) := by +theorem integralNormalization_coeff_degree_ne {i : ℕ} (hi : p.degree ≠ i) : + coeff (integralNormalization p) i = coeff p i * p.leadingCoeff ^ (p.natDegree - 1 - i) := by rw [integralNormalization_coeff, if_neg hi] -theorem integralNormalization_coeff_ne_natDegree {f : R[X]} {i : ℕ} (hi : i ≠ natDegree f) : - coeff (integralNormalization f) i = coeff f i * f.leadingCoeff ^ (f.natDegree - 1 - i) := - integralNormalization_coeff_ne_degree (degree_ne_of_natDegree_ne hi.symm) +theorem integralNormalization_coeff_ne_natDegree {i : ℕ} (hi : i ≠ natDegree p) : + coeff (integralNormalization p) i = coeff p i * p.leadingCoeff ^ (p.natDegree - 1 - i) := + integralNormalization_coeff_degree_ne (degree_ne_of_natDegree_ne hi.symm) -theorem monic_integralNormalization {f : R[X]} (hf : f ≠ 0) : Monic (integralNormalization f) := - monic_of_degree_le f.natDegree +theorem monic_integralNormalization (hp : p ≠ 0) : Monic (integralNormalization p) := + monic_of_degree_le p.natDegree (Finset.sup_le fun i h => - WithBot.coe_le_coe.2 <| le_natDegree_of_mem_supp i <| integralNormalization_support h) - (integralNormalization_coeff_natDegree hf) + WithBot.coe_le_coe.2 <| le_natDegree_of_mem_supp i <| support_integralNormalization_subset h) + (integralNormalization_coeff_natDegree hp) + +theorem integralNormalization_coeff_mul_leadingCoeff_pow (i : ℕ) (hp : 1 ≤ natDegree p) : + (integralNormalization p).coeff i * p.leadingCoeff ^ i = + p.coeff i * p.leadingCoeff ^ (p.natDegree - 1) := by + rw [integralNormalization_coeff] + split_ifs with h + · simp [natDegree_eq_of_degree_eq_some h, leadingCoeff, + ← pow_succ', tsub_add_cancel_of_le (natDegree_eq_of_degree_eq_some h ▸ hp)] + · simp only [mul_assoc, ← pow_add] + by_cases h' : i < p.degree + · rw [tsub_add_cancel_of_le] + rw [le_tsub_iff_right hp, Nat.succ_le_iff] + exact coe_lt_degree.mp h' + · simp [coeff_eq_zero_of_degree_lt (lt_of_le_of_ne (le_of_not_gt h') h)] + +theorem integralNormalization_mul_C_leadingCoeff (p : R[X]) : + integralNormalization p * C p.leadingCoeff = scaleRoots p p.leadingCoeff := by + ext i + rw [coeff_mul_C, integralNormalization_coeff] + split_ifs with h + · simp [natDegree_eq_of_degree_eq_some h, leadingCoeff] + · simp only [ge_iff_le, tsub_le_iff_right, smul_eq_mul, coeff_scaleRoots] + by_cases h' : i < p.degree + · rw [mul_assoc, ← pow_succ, tsub_right_comm, tsub_add_cancel_of_le] + rw [le_tsub_iff_left (coe_lt_degree.mp h').le, Nat.succ_le_iff] + exact coe_lt_degree.mp h' + · simp [coeff_eq_zero_of_degree_lt (lt_of_le_of_ne (le_of_not_gt h') h)] + +theorem integralNormalization_degree : (integralNormalization p).degree = p.degree := by + apply le_antisymm + · exact Finset.sup_mono p.support_integralNormalization_subset + · rw [← degree_scaleRoots, ← integralNormalization_mul_C_leadingCoeff] + exact (degree_mul_le _ _).trans (add_le_of_nonpos_right degree_C_le) + +variable [CommSemiring S] + +theorem leadingCoeff_smul_integralNormalization (p : S[X]) : + p.leadingCoeff • integralNormalization p = scaleRoots p p.leadingCoeff := by + rw [Algebra.smul_def, algebraMap_eq, mul_comm, integralNormalization_mul_C_leadingCoeff] + +theorem integralNormalization_eval₂_leadingCoeff_mul (h : 1 ≤ p.natDegree) (f : R →+* S) (x : S) : + (integralNormalization p).eval₂ f (f p.leadingCoeff * x) = + f p.leadingCoeff ^ (p.natDegree - 1) * p.eval₂ f x := by + rw [eval₂_eq_sum_range, eval₂_eq_sum_range, Finset.mul_sum] + apply Finset.sum_congr + · rw [natDegree_eq_of_degree_eq p.integralNormalization_degree] + intro n _hn + rw [mul_pow, ← mul_assoc, ← f.map_pow, ← f.map_mul, + integralNormalization_coeff_mul_leadingCoeff_pow _ h, f.map_mul, f.map_pow] + ring + +theorem integralNormalization_eval₂_eq_zero {p : R[X]} (f : R →+* S) {z : S} (hz : eval₂ f z p = 0) + (inj : ∀ x : R, f x = 0 → x = 0) : + eval₂ f (f p.leadingCoeff * z) (integralNormalization p) = 0 := by + obtain (h | h) := p.natDegree.eq_zero_or_pos + · by_cases h0 : coeff p 0 = 0 + · rw [eq_C_of_natDegree_eq_zero h] + simp [h0] + · rw [eq_C_of_natDegree_eq_zero h, eval₂_C] at hz + exact absurd (inj _ hz) h0 + · rw [integralNormalization_eval₂_leadingCoeff_mul h, hz, mul_zero] end Semiring -section IsDomain +section CommSemiring + +variable [CommSemiring R] [CommSemiring S] -variable [Ring R] [IsDomain R] +theorem integralNormalization_aeval_eq_zero [Algebra R S] {f : R[X]} {z : S} (hz : aeval z f = 0) + (inj : ∀ x : R, algebraMap R S x = 0 → x = 0) : + aeval (algebraMap R S f.leadingCoeff * z) (integralNormalization f) = 0 := + integralNormalization_eval₂_eq_zero (algebraMap R S) hz inj + +end CommSemiring + +section IsCancelMulZero + +variable [Semiring R] [IsCancelMulZero R] @[simp] theorem support_integralNormalization {f : R[X]} : (integralNormalization f).support = f.support := by + nontriviality R using Subsingleton.eq_zero + have : IsDomain R := {} by_cases hf : f = 0; · simp [hf] ext i - refine ⟨fun h => integralNormalization_support h, ?_⟩ + refine ⟨fun h => support_integralNormalization_subset h, ?_⟩ simp only [integralNormalization_coeff, mem_support_iff] intro hfi split_ifs with hi <;> simp [hf, hfi, hi] -end IsDomain - -section IsDomain - -variable [CommRing R] [IsDomain R] -variable [CommSemiring S] - -theorem integralNormalization_eval₂_eq_zero {p : R[X]} (f : R →+* S) {z : S} (hz : eval₂ f z p = 0) - (inj : ∀ x : R, f x = 0 → x = 0) : - eval₂ f (z * f p.leadingCoeff) (integralNormalization p) = 0 := - calc - eval₂ f (z * f p.leadingCoeff) (integralNormalization p) = - p.support.attach.sum fun i => - f (coeff (integralNormalization p) i.1 * p.leadingCoeff ^ i.1) * z ^ i.1 := by - rw [eval₂_eq_sum, sum_def, support_integralNormalization] - simp only [mul_comm z, mul_pow, mul_assoc, RingHom.map_pow, RingHom.map_mul] - rw [← Finset.sum_attach] - _ = - p.support.attach.sum fun i => - f (coeff p i.1 * p.leadingCoeff ^ (natDegree p - 1)) * z ^ i.1 := by - by_cases hp : p = 0; · simp [hp] - have one_le_deg : 1 ≤ natDegree p := - Nat.succ_le_of_lt (natDegree_pos_of_eval₂_root hp f hz inj) - congr with i - congr 2 - by_cases hi : i.1 = natDegree p - · rw [hi, integralNormalization_coeff_degree, one_mul, leadingCoeff, ← pow_succ', - tsub_add_cancel_of_le one_le_deg] - exact degree_eq_natDegree hp - · have : i.1 ≤ p.natDegree - 1 := - Nat.le_sub_one_of_lt - (lt_of_le_of_ne (le_natDegree_of_ne_zero (mem_support_iff.mp i.2)) hi) - rw [integralNormalization_coeff_ne_natDegree hi, mul_assoc, ← pow_add, - tsub_add_cancel_of_le this] - _ = f p.leadingCoeff ^ (natDegree p - 1) * eval₂ f z p := by - simp_rw [eval₂_eq_sum, sum_def, fun i => mul_comm (coeff p i), RingHom.map_mul, - RingHom.map_pow, mul_assoc, ← Finset.mul_sum] - congr 1 - exact p.support.sum_attach fun i ↦ f (p.coeff i) * z ^ i - _ = 0 := by rw [hz, mul_zero] - -theorem integralNormalization_aeval_eq_zero [Algebra R S] {f : R[X]} {z : S} (hz : aeval z f = 0) - (inj : ∀ x : R, algebraMap R S x = 0 → x = 0) : - aeval (z * algebraMap R S f.leadingCoeff) (integralNormalization f) = 0 := - integralNormalization_eval₂_eq_zero (algebraMap R S) hz inj - -end IsDomain +end IsCancelMulZero end IntegralNormalization From 4fafb9f5b76af48f96d2366beee76af2113161f7 Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Sat, 30 Nov 2024 19:56:26 +0000 Subject: [PATCH 435/829] feat(CategoryTheory): Relation between the Grothendieck construction and `AsSmall` (#19539) --- Mathlib.lean | 1 + .../CategoryTheory/Category/Cat/AsSmall.lean | 34 +++++++++ Mathlib/CategoryTheory/Category/ULift.lean | 10 +++ Mathlib/CategoryTheory/Grothendieck.lean | 71 ++++++++++++++++--- 4 files changed, 108 insertions(+), 8 deletions(-) create mode 100644 Mathlib/CategoryTheory/Category/Cat/AsSmall.lean diff --git a/Mathlib.lean b/Mathlib.lean index b2c20556f8317..e7bf4ac3e66d4 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1624,6 +1624,7 @@ import Mathlib.CategoryTheory.Category.Basic import Mathlib.CategoryTheory.Category.Bipointed import Mathlib.CategoryTheory.Category.Cat import Mathlib.CategoryTheory.Category.Cat.Adjunction +import Mathlib.CategoryTheory.Category.Cat.AsSmall import Mathlib.CategoryTheory.Category.Cat.Limit import Mathlib.CategoryTheory.Category.Factorisation import Mathlib.CategoryTheory.Category.GaloisConnection diff --git a/Mathlib/CategoryTheory/Category/Cat/AsSmall.lean b/Mathlib/CategoryTheory/Category/Cat/AsSmall.lean new file mode 100644 index 0000000000000..f65cd41c8bb6d --- /dev/null +++ b/Mathlib/CategoryTheory/Category/Cat/AsSmall.lean @@ -0,0 +1,34 @@ +/- +Copyright (c) 2024 Jakob von Raumer. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jakob von Raumer +-/ +import Mathlib.CategoryTheory.Category.Cat +import Mathlib.CategoryTheory.Category.ULift + +/-! +# Functorially embedding `Cat` into the category of small categories + +There is a canonical functor `asSmallFunctor` between the category of categories of any size and +any larger category of small categories. + +## Future Work + +Show that `asSmallFunctor` is faithful. +-/ + +universe w v u + +namespace CategoryTheory + +namespace Cat + +/-- Assigning to each category `C` the small category `AsSmall C` induces a functor `Cat ⥤ Cat`. -/ +@[simps] +def asSmallFunctor : Cat.{v, u} ⥤ Cat.{max w v u, max w v u} where + obj C := .of <| AsSmall C + map F := AsSmall.down ⋙ F ⋙ AsSmall.up + +end Cat + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Category/ULift.lean b/Mathlib/CategoryTheory/Category/ULift.lean index f70fe083dec2f..6a7b80e9fb347 100644 --- a/Mathlib/CategoryTheory/Category/ULift.lean +++ b/Mathlib/CategoryTheory/Category/ULift.lean @@ -174,6 +174,16 @@ def AsSmall.down : AsSmall C ⥤ C where obj X := ULift.down X map f := f.down +@[reassoc] +theorem down_comp {X Y Z : AsSmall C} (f : X ⟶ Y) (g : Y ⟶ Z) : (f ≫ g).down = f.down ≫ g.down := + rfl + +@[simp] +theorem eqToHom_down {X Y : AsSmall C} (h : X = Y) : + (eqToHom h).down = eqToHom (congrArg ULift.down h) := by + subst h + rfl + /-- The equivalence between `C` and `AsSmall C`. -/ @[simps] def AsSmall.equiv : C ≌ AsSmall C where diff --git a/Mathlib/CategoryTheory/Grothendieck.lean b/Mathlib/CategoryTheory/Grothendieck.lean index eeffded6b2bbb..78495117317b8 100644 --- a/Mathlib/CategoryTheory/Grothendieck.lean +++ b/Mathlib/CategoryTheory/Grothendieck.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison, Sina Hazratpour -/ -import Mathlib.CategoryTheory.Category.Cat +import Mathlib.CategoryTheory.Category.Cat.AsSmall import Mathlib.CategoryTheory.Elements import Mathlib.CategoryTheory.Comma.Over @@ -38,12 +38,13 @@ See also `CategoryTheory.Functor.Elements` for the category of elements of funct -/ -universe u +universe w u v u₁ v₁ u₂ v₂ namespace CategoryTheory -variable {C D : Type*} [Category C] [Category D] -variable (F : C ⥤ Cat) +variable {C : Type u} [Category.{v} C] +variable {D : Type u₁} [Category.{v₁} D] +variable (F : C ⥤ Cat.{v₂, u₂}) /-- The Grothendieck construction (often written as `∫ F` in mathematics) for a functor `F : C ⥤ Cat` @@ -229,9 +230,65 @@ theorem map_comp_eq (α : F ⟶ G) (β : G ⟶ H) : if possible, and we should prefer `map_comp_iso` to `map_comp_eq` whenever we can. -/ def mapCompIso (α : F ⟶ G) (β : G ⟶ H) : map (α ≫ β) ≅ map α ⋙ map β := eqToIso (map_comp_eq α β) -end +variable (F) -universe v +/-- The inverse functor to build the equivalence `compAsSmallFunctorEquivalence`. -/ +@[simps] +def compAsSmallFunctorEquivalenceInverse : + Grothendieck F ⥤ Grothendieck (F ⋙ Cat.asSmallFunctor.{w}) where + obj X := ⟨X.base, AsSmall.up.obj X.fiber⟩ + map f := ⟨f.base, AsSmall.up.map f.fiber⟩ + +/-- The functor to build the equivalence `compAsSmallFunctorEquivalence`. -/ +@[simps] +def compAsSmallFunctorEquivalenceFunctor : + Grothendieck (F ⋙ Cat.asSmallFunctor.{w}) ⥤ Grothendieck F where + obj X := ⟨X.base, AsSmall.down.obj X.fiber⟩ + map f := ⟨f.base, AsSmall.down.map f.fiber⟩ + map_id _ := by apply Grothendieck.ext <;> simp + map_comp _ _ := by apply Grothendieck.ext <;> simp [down_comp] + +/-- Taking the Grothendieck construction on `F ⋙ asSmallFunctor`, where +`asSmallFunctor : Cat ⥤ Cat` is the functor which turns each category into a small category of a +(potentiall) larger universe, is equivalent to the Grothendieck construction on `F` itself. -/ +@[simps] +def compAsSmallFunctorEquivalence : + Grothendieck (F ⋙ Cat.asSmallFunctor.{w}) ≌ Grothendieck F where + functor := compAsSmallFunctorEquivalenceFunctor F + inverse := compAsSmallFunctorEquivalenceInverse F + counitIso := Iso.refl _ + unitIso := Iso.refl _ + +/-- Mapping a Grothendieck construction along the whiskering of any natural transformation +`α : F ⟶ G` with the functor `asSmallFunctor : Cat ⥤ Cat` is naturally isomorphic to conjugating +`map α` with the equivalence between `Grothendieck (F ⋙ asSmallFunctor)` and `Grothendieck F`. -/ +def mapWhiskerRightAsSmallFunctor (α : F ⟶ G) : + map (whiskerRight α Cat.asSmallFunctor.{w}) ≅ + (compAsSmallFunctorEquivalence F).functor ⋙ map α ⋙ + (compAsSmallFunctorEquivalence G).inverse := + NatIso.ofComponents + (fun X => Iso.refl _) + (fun f => by + fapply Grothendieck.ext + · simp [compAsSmallFunctorEquivalenceInverse] + · simp only [compAsSmallFunctorEquivalence_functor, compAsSmallFunctorEquivalence_inverse, + Functor.comp_obj, compAsSmallFunctorEquivalenceInverse_obj_base, map_obj_base, + compAsSmallFunctorEquivalenceFunctor_obj_base, Cat.asSmallFunctor_obj, Cat.of_α, + Iso.refl_hom, Functor.comp_map, comp_base, id_base, + compAsSmallFunctorEquivalenceInverse_map_base, map_map_base, + compAsSmallFunctorEquivalenceFunctor_map_base, Cat.asSmallFunctor_map, map_obj_fiber, + whiskerRight_app, AsSmall.down_obj, AsSmall.up_obj_down, + compAsSmallFunctorEquivalenceInverse_obj_fiber, + compAsSmallFunctorEquivalenceFunctor_obj_fiber, comp_fiber, map_map_fiber, + AsSmall.down_map, down_comp, eqToHom_down, AsSmall.up_map_down, Functor.map_comp, + eqToHom_map, id_fiber, Category.assoc, eqToHom_trans_assoc, + compAsSmallFunctorEquivalenceInverse_map_fiber, + compAsSmallFunctorEquivalenceFunctor_map_fiber, eqToHom_comp_iff, comp_eqToHom_iff] + simp only [eqToHom_trans_assoc, Category.assoc, conj_eqToHom_iff_heq'] + rw [G.map_id] + simp ) + +end /-- The Grothendieck construction as a functor from the functor category `E ⥤ Cat` to the over category `Over E`. -/ @@ -245,8 +302,6 @@ def functor {E : Cat.{v,u}} : (E ⥤ Cat.{v,u}) ⥤ Over (T := Cat.{v,u}) E wher simp [Grothendieck.map_comp_eq α β] rfl -universe w - variable (G : C ⥤ Type w) /-- Auxiliary definition for `grothendieckTypeToCat`, to speed up elaboration. -/ From 6c4358eaffb6961ea98200da8aaa4d063cae2d86 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 30 Nov 2024 21:42:31 +0000 Subject: [PATCH 436/829] chore: don't import conformal maps in Behrend's construction (#19638) --- Mathlib/Analysis/Complex/AbsMax.lean | 2 +- Mathlib/Analysis/Complex/CauchyIntegral.lean | 1 + Mathlib/Analysis/Complex/Conformal.lean | 53 +++++++++++++++- Mathlib/Analysis/Complex/RealDeriv.lean | 62 +------------------ .../SpecialFunctions/Complex/LogDeriv.lean | 2 + .../Analysis/SpecialFunctions/ExpDeriv.lean | 7 ++- .../Additive/AP/Three/Behrend.lean | 3 + .../Geometry/Manifold/Instances/Sphere.lean | 2 + .../Measure/Lebesgue/VolumeOfBalls.lean | 3 +- Mathlib/NumberTheory/EulerProduct/ExpLog.lean | 1 + .../ModularForms/JacobiTheta/TwoVariable.lean | 3 +- 11 files changed, 74 insertions(+), 65 deletions(-) diff --git a/Mathlib/Analysis/Complex/AbsMax.lean b/Mathlib/Analysis/Complex/AbsMax.lean index 9f5038937c5e6..d2af8a302da8d 100644 --- a/Mathlib/Analysis/Complex/AbsMax.lean +++ b/Mathlib/Analysis/Complex/AbsMax.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Analysis.Complex.CauchyIntegral -import Mathlib.Analysis.Normed.Module.Completion import Mathlib.Analysis.NormedSpace.Extr +import Mathlib.Data.Complex.FiniteDimensional import Mathlib.Topology.Order.ExtrClosure /-! diff --git a/Mathlib/Analysis/Complex/CauchyIntegral.lean b/Mathlib/Analysis/Complex/CauchyIntegral.lean index 340b252b8b626..19be8544d83fe 100644 --- a/Mathlib/Analysis/Complex/CauchyIntegral.lean +++ b/Mathlib/Analysis/Complex/CauchyIntegral.lean @@ -8,6 +8,7 @@ import Mathlib.Analysis.Calculus.DiffContOnCl import Mathlib.Analysis.Calculus.DSlope import Mathlib.Analysis.Calculus.FDeriv.Analytic import Mathlib.Analysis.Complex.ReImTopology +import Mathlib.Data.Real.Cardinality import Mathlib.MeasureTheory.Integral.CircleIntegral import Mathlib.MeasureTheory.Integral.DivergenceTheorem import Mathlib.MeasureTheory.Measure.Lebesgue.Complex diff --git a/Mathlib/Analysis/Complex/Conformal.lean b/Mathlib/Analysis/Complex/Conformal.lean index 18618155e7d9d..87d4b01c9333d 100644 --- a/Mathlib/Analysis/Complex/Conformal.lean +++ b/Mathlib/Analysis/Complex/Conformal.lean @@ -3,8 +3,11 @@ Copyright (c) 2021 Yourong Zang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yourong Zang -/ +import Mathlib.Analysis.Calculus.Conformal.NormedSpace +import Mathlib.Analysis.Calculus.Deriv.Basic +import Mathlib.Analysis.Calculus.FDeriv.Equiv +import Mathlib.Analysis.Calculus.FDeriv.RestrictScalars import Mathlib.Analysis.Complex.Isometry -import Mathlib.Analysis.NormedSpace.ConformalLinearMap import Mathlib.Analysis.Normed.Module.FiniteDimension import Mathlib.Data.Complex.FiniteDimensional @@ -25,10 +28,24 @@ to be conformal. linear or the composition of some complex linear map and `conj`. +* `DifferentiableAt.conformalAt` states that a real-differentiable function with a nonvanishing + differential from the complex plane into an arbitrary complex-normed space is conformal at a point + if it's holomorphic at that point. This is a version of Cauchy-Riemann equations. + +* `conformalAt_iff_differentiableAt_or_differentiableAt_comp_conj` proves that a real-differential + function with a nonvanishing differential between the complex plane is conformal at a point if and + only if it's holomorphic or antiholomorphic at that point. + ## Warning Antiholomorphic functions such as the complex conjugate are considered as conformal functions in this file. + +## TODO + +* The classical form of Cauchy-Riemann equations +* On a connected open set `u`, a function which is `ConformalAt` each point is either holomorphic +throughout or antiholomorphic throughout. -/ @@ -109,3 +126,37 @@ theorem isConformalMap_iff_is_complex_or_conj_linear : simp only [w, restrictScalars_zero, zero_comp] end ConformalIntoComplexPlane + +/-! ### Conformality of real-differentiable complex maps -/ + +section Conformality +variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℂ E] {z : ℂ} {f : ℂ → E} + +/-- A real differentiable function of the complex plane into some complex normed space `E` is +conformal at a point `z` if it is holomorphic at that point with a nonvanishing differential. +This is a version of the Cauchy-Riemann equations. -/ +theorem DifferentiableAt.conformalAt (h : DifferentiableAt ℂ f z) (hf' : deriv f z ≠ 0) : + ConformalAt f z := by + rw [conformalAt_iff_isConformalMap_fderiv, (h.hasFDerivAt.restrictScalars ℝ).fderiv] + apply isConformalMap_complex_linear + simpa only [Ne, ContinuousLinearMap.ext_ring_iff] + +/-- A complex function is conformal if and only if the function is holomorphic or antiholomorphic +with a nonvanishing differential. -/ +theorem conformalAt_iff_differentiableAt_or_differentiableAt_comp_conj {f : ℂ → ℂ} {z : ℂ} : + ConformalAt f z ↔ + (DifferentiableAt ℂ f z ∨ DifferentiableAt ℂ (f ∘ conj) (conj z)) ∧ fderiv ℝ f z ≠ 0 := by + rw [conformalAt_iff_isConformalMap_fderiv] + rw [isConformalMap_iff_is_complex_or_conj_linear] + apply and_congr_left + intro h + have h_diff := h.imp_symm fderiv_zero_of_not_differentiableAt + apply or_congr + · rw [differentiableAt_iff_restrictScalars ℝ h_diff] + rw [← conj_conj z] at h_diff + rw [differentiableAt_iff_restrictScalars ℝ (h_diff.comp _ conjCLE.differentiableAt)] + refine exists_congr fun g => rfl.congr ?_ + have : fderiv ℝ conj (conj z) = _ := conjCLE.fderiv + simp [fderiv_comp _ h_diff conjCLE.differentiableAt, this, conj_conj] + +end Conformality diff --git a/Mathlib/Analysis/Complex/RealDeriv.lean b/Mathlib/Analysis/Complex/RealDeriv.lean index 1609fa5e357a4..92f1c438dbbbb 100644 --- a/Mathlib/Analysis/Complex/RealDeriv.lean +++ b/Mathlib/Analysis/Complex/RealDeriv.lean @@ -5,34 +5,17 @@ Authors: Sébastien Gouëzel, Yourong Zang -/ import Mathlib.Analysis.Calculus.ContDiff.Basic import Mathlib.Analysis.Calculus.Deriv.Linear -import Mathlib.Analysis.Complex.Conformal -import Mathlib.Analysis.Calculus.Conformal.NormedSpace +import Mathlib.Analysis.Complex.Basic /-! # Real differentiability of complex-differentiable functions `HasDerivAt.real_of_complex` expresses that, if a function on `ℂ` is differentiable (over `ℂ`), then its restriction to `ℝ` is differentiable over `ℝ`, with derivative the real part of the complex derivative. - -`DifferentiableAt.conformalAt` states that a real-differentiable function with a nonvanishing -differential from the complex plane into an arbitrary complex-normed space is conformal at a point -if it's holomorphic at that point. This is a version of Cauchy-Riemann equations. - -`conformalAt_iff_differentiableAt_or_differentiableAt_comp_conj` proves that a real-differential -function with a nonvanishing differential between the complex plane is conformal at a point if and -only if it's holomorphic or antiholomorphic at that point. - -## TODO - -* The classical form of Cauchy-Riemann equations -* On a connected open set `u`, a function which is `ConformalAt` each point is either holomorphic -throughout or antiholomorphic throughout. - -## Warning - -We do NOT require conformal functions to be orientation-preserving in this file. -/ +assert_not_exists IsConformalMap +assert_not_exists Conformal section RealDerivOfComplex @@ -131,42 +114,3 @@ theorem HasDerivAt.ofReal_comp {f : ℝ → ℝ} {u : ℝ} (hf : HasDerivAt f u ofRealCLM.hasDerivAt.scomp z hf end RealDerivOfComplex - -section Conformality - -/-! ### Conformality of real-differentiable complex maps -/ - -open Complex ContinuousLinearMap - -open scoped ComplexConjugate - -variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℂ E] {z : ℂ} {f : ℂ → E} - -/-- A real differentiable function of the complex plane into some complex normed space `E` is - conformal at a point `z` if it is holomorphic at that point with a nonvanishing differential. - This is a version of the Cauchy-Riemann equations. -/ -theorem DifferentiableAt.conformalAt (h : DifferentiableAt ℂ f z) (hf' : deriv f z ≠ 0) : - ConformalAt f z := by - rw [conformalAt_iff_isConformalMap_fderiv, (h.hasFDerivAt.restrictScalars ℝ).fderiv] - apply isConformalMap_complex_linear - simpa only [Ne, ContinuousLinearMap.ext_ring_iff] - -/-- A complex function is conformal if and only if the function is holomorphic or antiholomorphic - with a nonvanishing differential. -/ -theorem conformalAt_iff_differentiableAt_or_differentiableAt_comp_conj {f : ℂ → ℂ} {z : ℂ} : - ConformalAt f z ↔ - (DifferentiableAt ℂ f z ∨ DifferentiableAt ℂ (f ∘ conj) (conj z)) ∧ fderiv ℝ f z ≠ 0 := by - rw [conformalAt_iff_isConformalMap_fderiv] - rw [isConformalMap_iff_is_complex_or_conj_linear] - apply and_congr_left - intro h - have h_diff := h.imp_symm fderiv_zero_of_not_differentiableAt - apply or_congr - · rw [differentiableAt_iff_restrictScalars ℝ h_diff] - rw [← conj_conj z] at h_diff - rw [differentiableAt_iff_restrictScalars ℝ (h_diff.comp _ conjCLE.differentiableAt)] - refine exists_congr fun g => rfl.congr ?_ - have : fderiv ℝ conj (conj z) = _ := conjCLE.fderiv - simp [fderiv_comp _ h_diff conjCLE.differentiableAt, this, conj_conj] - -end Conformality diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/LogDeriv.lean b/Mathlib/Analysis/SpecialFunctions/Complex/LogDeriv.lean index 305f06d9e1523..8274b1696a1db 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/LogDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/LogDeriv.lean @@ -13,6 +13,8 @@ import Mathlib.Analysis.SpecialFunctions.ExpDeriv -/ +assert_not_exists IsConformalMap +assert_not_exists Conformal open Set Filter diff --git a/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean b/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean index 805c933b00240..6a54e8ef68876 100644 --- a/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean @@ -3,10 +3,10 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. 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.FDeriv.Analytic import Mathlib.Analysis.Calculus.IteratedDeriv.Lemmas +import Mathlib.Analysis.Complex.RealDeriv +import Mathlib.Analysis.SpecialFunctions.Exp import Mathlib.Analysis.SpecialFunctions.Exponential /-! @@ -19,6 +19,9 @@ In this file we prove that `Complex.exp` and `Real.exp` are infinitely smooth fu exp, derivative -/ +assert_not_exists IsConformalMap +assert_not_exists Conformal + noncomputable section open Filter Asymptotics Set Function diff --git a/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean b/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean index 19ecd990580b6..2dd42b91269fb 100644 --- a/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean +++ b/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean @@ -42,6 +42,9 @@ integer points on that sphere and map them onto `ℕ` in a way that preserves ar 3AP-free, Salem-Spencer, Behrend construction, arithmetic progression, sphere, strictly convex -/ +assert_not_exists IsConformalMap +assert_not_exists Conformal + open Nat hiding log open Finset Metric Real open scoped Pointwise diff --git a/Mathlib/Geometry/Manifold/Instances/Sphere.lean b/Mathlib/Geometry/Manifold/Instances/Sphere.lean index a426774e69c58..abc2d5450eb3b 100644 --- a/Mathlib/Geometry/Manifold/Instances/Sphere.lean +++ b/Mathlib/Geometry/Manifold/Instances/Sphere.lean @@ -4,10 +4,12 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Heather Macbeth -/ import Mathlib.Analysis.Calculus.Deriv.Inv +import Mathlib.Analysis.Complex.Circle import Mathlib.Analysis.NormedSpace.BallAction import Mathlib.Analysis.SpecialFunctions.ExpDeriv import Mathlib.Analysis.InnerProductSpace.Calculus import Mathlib.Analysis.InnerProductSpace.PiL2 +import Mathlib.Data.Complex.FiniteDimensional import Mathlib.Geometry.Manifold.Algebra.LieGroup import Mathlib.Geometry.Manifold.Instances.Real import Mathlib.Geometry.Manifold.MFDeriv.Basic diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean index be988e8efe5a8..ad3c3db1e4080 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean @@ -3,10 +3,11 @@ Copyright (c) 2023 Xavier Roblot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Xavier Roblot -/ +import Mathlib.Analysis.SpecialFunctions.Gamma.BohrMollerup +import Mathlib.Data.Complex.FiniteDimensional import Mathlib.MeasureTheory.Constructions.HaarToSphere import Mathlib.MeasureTheory.Integral.Gamma import Mathlib.MeasureTheory.Integral.Pi -import Mathlib.Analysis.SpecialFunctions.Gamma.BohrMollerup /-! # Volume of balls diff --git a/Mathlib/NumberTheory/EulerProduct/ExpLog.lean b/Mathlib/NumberTheory/EulerProduct/ExpLog.lean index 8826ef2f3eba6..eb2c16084793f 100644 --- a/Mathlib/NumberTheory/EulerProduct/ExpLog.lean +++ b/Mathlib/NumberTheory/EulerProduct/ExpLog.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Stoll -/ import Mathlib.Analysis.SpecialFunctions.Complex.LogDeriv +import Mathlib.Data.Complex.FiniteDimensional import Mathlib.NumberTheory.EulerProduct.Basic /-! diff --git a/Mathlib/NumberTheory/ModularForms/JacobiTheta/TwoVariable.lean b/Mathlib/NumberTheory/ModularForms/JacobiTheta/TwoVariable.lean index ff8f8efa05056..a8e8dfc0a7835 100644 --- a/Mathlib/NumberTheory/ModularForms/JacobiTheta/TwoVariable.lean +++ b/Mathlib/NumberTheory/ModularForms/JacobiTheta/TwoVariable.lean @@ -3,9 +3,10 @@ Copyright (c) 2023 David Loeffler. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Loeffler -/ -import Mathlib.Analysis.SpecialFunctions.Gaussian.PoissonSummation import Mathlib.Analysis.Calculus.SmoothSeries import Mathlib.Analysis.NormedSpace.OperatorNorm.Prod +import Mathlib.Analysis.SpecialFunctions.Gaussian.PoissonSummation +import Mathlib.Data.Complex.FiniteDimensional /-! # The two-variable Jacobi theta function From fa4bdc153de73cc43e80ce0cee6a638e45110a33 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sun, 1 Dec 2024 00:08:39 +0000 Subject: [PATCH 437/829] feat: introduce ENat.card, and use it in place of PartENat.card (#19624) --- Mathlib/Algebra/Group/Pointwise/Set/Card.lean | 4 +- Mathlib/Analysis/Convex/Radon.lean | 4 +- Mathlib/Data/Finite/Card.lean | 12 ++ Mathlib/Data/Set/Card.lean | 29 ++--- Mathlib/SetTheory/Cardinal/ENat.lean | 19 +++ Mathlib/SetTheory/Cardinal/Finite.lean | 117 ++++++++++++++++-- 6 files changed, 152 insertions(+), 33 deletions(-) diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean index 812782481817b..d6ccb0c30a835 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean @@ -62,7 +62,7 @@ attribute [deprecated natCard_inv (since := "2024-09-30")] card_inv attribute [deprecated natCard_neg (since := "2024-09-30")] card_neg @[to_additive (attr := simp)] -lemma encard_inv (s : Set G) : s⁻¹.encard = s.encard := by simp [encard, PartENat.card] +lemma encard_inv (s : Set G) : s⁻¹.encard = s.encard := by simp [encard, ENat.card] @[to_additive (attr := simp)] lemma ncard_inv (s : Set G) : s⁻¹.ncard = s.ncard := by simp [ncard] @@ -112,7 +112,7 @@ attribute [deprecated Cardinal.mk_vadd_set (since := "2024-09-30")] card_vadd_se @[to_additive (attr := simp)] lemma encard_smul_set (a : G) (s : Set α) : (a • s).encard = s.encard := by - simp [encard, PartENat.card] + simp [encard, ENat.card] @[to_additive (attr := simp)] lemma ncard_smul_set (a : G) (s : Set α) : (a • s).ncard = s.ncard := by simp [ncard] diff --git a/Mathlib/Analysis/Convex/Radon.lean b/Mathlib/Analysis/Convex/Radon.lean index 21470d3e027c9..ae64b4a1822a7 100644 --- a/Mathlib/Analysis/Convex/Radon.lean +++ b/Mathlib/Analysis/Convex/Radon.lean @@ -212,7 +212,7 @@ If `F` is a (possibly infinite) family of more than `d + 1` compact convex sets finite dimension `d`, and any `d + 1` sets of `F` intersect nontrivially, then all sets of `F` intersect nontrivially. -/ theorem helly_theorem_compact [TopologicalSpace E] [T2Space E] {F : ι → Set E} - (h_card : finrank 𝕜 E + 1 ≤ PartENat.card ι) + (h_card : finrank 𝕜 E + 1 ≤ ENat.card ι) (h_convex : ∀ i, Convex 𝕜 (F i)) (h_compact : ∀ i, IsCompact (F i)) (h_inter : ∀ I : Finset ι, #I = finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : (⋂ i, F i).Nonempty := by @@ -224,7 +224,7 @@ theorem helly_theorem_compact [TopologicalSpace E] [T2Space E] {F : ι → Set E · have : Finite ι := Finite.of_not_infinite h have : Fintype ι := Fintype.ofFinite ι apply exists_superset_card_eq hI_card - simp only [PartENat.card_eq_coe_fintype_card] at h_card + simp only [ENat.card_eq_coe_fintype_card] at h_card rwa [← Nat.cast_one, ← Nat.cast_add, Nat.cast_le] at h_card obtain ⟨J, hJ_ss, hJ_card⟩ := hJ apply Set.Nonempty.mono <| biInter_mono hJ_ss (by intro _ _; rfl) diff --git a/Mathlib/Data/Finite/Card.lean b/Mathlib/Data/Finite/Card.lean index fb6afe0f1c2f9..4f038d6ecdef9 100644 --- a/Mathlib/Data/Finite/Card.lean +++ b/Mathlib/Data/Finite/Card.lean @@ -162,8 +162,20 @@ theorem card_subtype_lt [Finite α] {p : α → Prop} {x : α} (hx : ¬p x) : end Finite +namespace ENat + +theorem card_eq_coe_natCard (α : Type*) [Finite α] : card α = Nat.card α := by + unfold ENat.card + apply symm + rw [Cardinal.natCast_eq_toENat_iff] + exact Finite.cast_card_eq_mk + +end ENat + namespace PartENat +set_option linter.deprecated false in +@[deprecated ENat.card_eq_coe_natCard (since := "2024-11-30")] theorem card_eq_coe_natCard (α : Type*) [Finite α] : card α = Nat.card α := by unfold PartENat.card apply symm diff --git a/Mathlib/Data/Set/Card.lean b/Mathlib/Data/Set/Card.lean index e6992277d23f0..4588229a6887f 100644 --- a/Mathlib/Data/Set/Card.lean +++ b/Mathlib/Data/Set/Card.lean @@ -10,7 +10,7 @@ import Mathlib.SetTheory.Cardinal.Finite We define the cardinality of set `s` as a term `Set.encard s : ℕ∞` and a term `Set.ncard s : ℕ`. The latter takes the junk value of zero if `s` is infinite. Both functions are noncomputable, and -are defined in terms of `PartENat.card` (which takes a type as its argument); this file can be seen +are defined in terms of `ENat.card` (which takes a type as its argument); this file can be seen as an API for the same function in the special case where the type is a coercion of a `Set`, allowing for smoother interactions with the `Set` API. @@ -59,19 +59,18 @@ namespace Set variable {α β : Type*} {s t : Set α} /-- The cardinality of a set as a term in `ℕ∞` -/ -noncomputable def encard (s : Set α) : ℕ∞ := PartENat.withTopEquiv (PartENat.card s) +noncomputable def encard (s : Set α) : ℕ∞ := ENat.card s @[simp] theorem encard_univ_coe (s : Set α) : encard (univ : Set s) = encard s := by - rw [encard, encard, PartENat.card_congr (Equiv.Set.univ ↑s)] + rw [encard, encard, ENat.card_congr (Equiv.Set.univ ↑s)] theorem encard_univ (α : Type*) : - encard (univ : Set α) = PartENat.withTopEquiv (PartENat.card α) := by - rw [encard, PartENat.card_congr (Equiv.Set.univ α)] + encard (univ : Set α) = ENat.card α := by + rw [encard, ENat.card_congr (Equiv.Set.univ α)] theorem Finite.encard_eq_coe_toFinset_card (h : s.Finite) : s.encard = h.toFinset.card := by have := h.fintype - rw [encard, PartENat.card_eq_coe_fintype_card, - PartENat.withTopEquiv_natCast, toFinite_toFinset, toFinset_card] + rw [encard, ENat.card_eq_coe_fintype_card, toFinite_toFinset, toFinset_card] theorem encard_eq_coe_toFinset_card (s : Set α) [Fintype s] : encard s = s.toFinset.card := by have h := toFinite s @@ -83,13 +82,10 @@ theorem encard_eq_coe_toFinset_card (s : Set α) [Fintype s] : encard s = s.toFi theorem Infinite.encard_eq {s : Set α} (h : s.Infinite) : s.encard = ⊤ := by have := h.to_subtype - rw [encard, ← PartENat.withTopEquiv.symm.injective.eq_iff, Equiv.symm_apply_apply, - PartENat.withTopEquiv_symm_top, PartENat.card_eq_top_of_infinite] + rw [encard, ENat.card_eq_top_of_infinite] @[simp] theorem encard_eq_zero : s.encard = 0 ↔ s = ∅ := by - rw [encard, ← PartENat.withTopEquiv.symm.injective.eq_iff, Equiv.symm_apply_apply, - PartENat.withTopEquiv_symm_zero, PartENat.card_eq_zero_iff_empty, isEmpty_subtype, - eq_empty_iff_forall_not_mem] + rw [encard, ENat.card_eq_zero_iff_empty, isEmpty_subtype, eq_empty_iff_forall_not_mem] @[simp] theorem encard_empty : (∅ : Set α).encard = 0 := by rw [encard_eq_zero] @@ -106,12 +102,11 @@ theorem encard_ne_zero : s.encard ≠ 0 ↔ s.Nonempty := by protected alias ⟨_, Nonempty.encard_pos⟩ := encard_pos @[simp] theorem encard_singleton (e : α) : ({e} : Set α).encard = 1 := by - rw [encard, ← PartENat.withTopEquiv.symm.injective.eq_iff, Equiv.symm_apply_apply, - PartENat.card_eq_coe_fintype_card, Fintype.card_ofSubsingleton, Nat.cast_one]; rfl + rw [encard, ENat.card_eq_coe_fintype_card, Fintype.card_ofSubsingleton, Nat.cast_one] theorem encard_union_eq (h : Disjoint s t) : (s ∪ t).encard = s.encard + t.encard := by classical - simp [encard, PartENat.card_congr (Equiv.Set.union h), PartENat.card_sum, PartENat.withTopEquiv] + simp [encard, ENat.card_congr (Equiv.Set.union h)] theorem encard_insert_of_not_mem {a : α} (has : a ∉ s) : (insert a s).encard = s.encard + 1 := by rw [← union_singleton, encard_union_eq (by simpa), encard_singleton] @@ -381,10 +376,10 @@ section Function variable {s : Set α} {t : Set β} {f : α → β} theorem InjOn.encard_image (h : InjOn f s) : (f '' s).encard = s.encard := by - rw [encard, PartENat.card_image_of_injOn h, encard] + rw [encard, ENat.card_image_of_injOn h, encard] theorem encard_congr (e : s ≃ t) : s.encard = t.encard := by - rw [← encard_univ_coe, ← encard_univ_coe t, encard_univ, encard_univ, PartENat.card_congr e] + rw [← encard_univ_coe, ← encard_univ_coe t, encard_univ, encard_univ, ENat.card_congr e] theorem _root_.Function.Injective.encard_image (hf : f.Injective) (s : Set α) : (f '' s).encard = s.encard := diff --git a/Mathlib/SetTheory/Cardinal/ENat.lean b/Mathlib/SetTheory/Cardinal/ENat.lean index 07957687d320d..b444b8fb5cd1f 100644 --- a/Mathlib/SetTheory/Cardinal/ENat.lean +++ b/Mathlib/SetTheory/Cardinal/ENat.lean @@ -260,6 +260,23 @@ theorem toENat_lift {a : Cardinal.{v}} : toENat (lift.{u} a) = toENat a := by | inl ha => lift a to ℕ∞ using ha; simp | inr ha => simp [toENat_eq_top.2, ha] +theorem toENat_congr {α : Type u} {β : Type v} (e : α ≃ β) : toENat #α = toENat #β := by + rw [← toENat_lift, lift_mk_eq.{_, _,v}.mpr ⟨e⟩, toENat_lift] + +lemma toENat_le_iff_of_le_aleph0 {c c' : Cardinal} (h : c ≤ ℵ₀) : + toENat c ≤ toENat c' ↔ c ≤ c' := by + lift c to ℕ∞ using h + simp_rw [toENat_ofENat, enat_gc _] + +lemma toENat_le_iff_of_lt_aleph0 {c c' : Cardinal} (hc' : c' < ℵ₀) : + toENat c ≤ toENat c' ↔ c ≤ c' := by + lift c' to ℕ using hc' + simp_rw [toENat_nat, ← toENat_le_nat] + +lemma toENat_eq_iff_of_le_aleph0 {c c' : Cardinal} (hc : c ≤ ℵ₀) (hc' : c' ≤ ℵ₀) : + toENat c = toENat c' ↔ c = c' := + toENat_strictMonoOn.injOn.eq_iff hc hc' + @[simp, norm_cast] lemma ofENat_add (m n : ℕ∞) : ofENat (m + n) = m + n := by apply toENat_injOn <;> simp @@ -288,4 +305,6 @@ def ofENatHom : ℕ∞ →+*o Cardinal where map_add' := ofENat_add monotone' := ofENat_mono + + end Cardinal diff --git a/Mathlib/SetTheory/Cardinal/Finite.lean b/Mathlib/SetTheory/Cardinal/Finite.lean index e31452e7c7598..6498c5e18a81a 100644 --- a/Mathlib/SetTheory/Cardinal/Finite.lean +++ b/Mathlib/SetTheory/Cardinal/Finite.lean @@ -14,8 +14,10 @@ import Mathlib.SetTheory.Cardinal.PartENat * `Nat.card α` is the cardinality of `α` as a natural number. If `α` is infinite, `Nat.card α = 0`. +* `ENat.card α` is the cardinality of `α` as an extended natural number. + If `α` is infinite, `ENat.card α = ⊤`. * `PartENat.card α` is the cardinality of `α` as an extended natural number - (using `Part ℕ`). If `α` is infinite, `PartENat.card α = ⊤`. + (using the legacy definition `PartENat := Part ℕ`). If `α` is infinite, `PartENat.card α = ⊤`. -/ open Cardinal Function @@ -242,85 +244,176 @@ protected alias ⟨_, Nonempty.natCard_pos⟩ := natCard_pos end Set + +namespace ENat + +/-- `ENat.card α` is the cardinality of `α` as an extended natural number. + If `α` is infinite, `ENat.card α = ⊤`. -/ +def card (α : Type*) : ℕ∞ := + toENat (mk α) + +@[simp] +theorem card_eq_coe_fintype_card [Fintype α] : card α = Fintype.card α := by + simp [card] + +@[simp] +theorem card_eq_top_of_infinite [Infinite α] : card α = ⊤ := by + simp [card] + +@[simp] +theorem card_sum (α β : Type*) : + card (α ⊕ β) = card α + card β := by + simp only [card, mk_sum, map_add, toENat_lift] + +theorem card_congr {α β : Type*} (f : α ≃ β) : card α = card β := + Cardinal.toENat_congr f + +@[simp] lemma card_ulift (α : Type*) : card (ULift α) = card α := card_congr Equiv.ulift + +@[simp] lemma card_plift (α : Type*) : card (PLift α) = card α := card_congr Equiv.plift + +theorem card_image_of_injOn {α β : Type*} {f : α → β} {s : Set α} (h : Set.InjOn f s) : + card (f '' s) = card s := + card_congr (Equiv.Set.imageOfInjOn f s h).symm + +theorem card_image_of_injective {α β : Type*} (f : α → β) (s : Set α) + (h : Function.Injective f) : card (f '' s) = card s := card_image_of_injOn h.injOn + +@[simp] +theorem _root_.Cardinal.natCast_le_toENat_iff {n : ℕ} {c : Cardinal} : + ↑n ≤ toENat c ↔ ↑n ≤ c := by + rw [← toENat_nat n, toENat_le_iff_of_le_aleph0 (le_of_lt (nat_lt_aleph0 n))] + +theorem _root_.Cardinal.toENat_le_natCast_iff {c : Cardinal} {n : ℕ} : + toENat c ≤ n ↔ c ≤ n := by simp + +@[simp] +theorem _root_.Cardinal.natCast_eq_toENat_iff {n : ℕ} {c : Cardinal} : + ↑n = toENat c ↔ ↑n = c := by + rw [le_antisymm_iff, le_antisymm_iff, Cardinal.toENat_le_natCast_iff, + Cardinal.natCast_le_toENat_iff] + +theorem _root_.Cardinal.toENat_eq_natCast_iff {c : Cardinal} {n : ℕ} : + Cardinal.toENat c = n ↔ c = n := by simp + +@[simp] +theorem _root_.Cardinal.natCast_lt_toENat_iff {n : ℕ} {c : Cardinal} : + ↑n < toENat c ↔ ↑n < c := by + simp only [← not_le, Cardinal.toENat_le_natCast_iff] + +@[simp] +theorem _root_.Cardinal.toENat_lt_natCast_iff {n : ℕ} {c : Cardinal} : + toENat c < ↑n ↔ c < ↑n := by + simp only [← not_le, Cardinal.natCast_le_toENat_iff] + +theorem card_eq_zero_iff_empty (α : Type*) : card α = 0 ↔ IsEmpty α := by + rw [← Cardinal.mk_eq_zero_iff] + simp [card] + +theorem card_le_one_iff_subsingleton (α : Type*) : card α ≤ 1 ↔ Subsingleton α := by + rw [← le_one_iff_subsingleton] + simp [card] + +theorem one_lt_card_iff_nontrivial (α : Type*) : 1 < card α ↔ Nontrivial α := by + rw [← Cardinal.one_lt_iff_nontrivial] + conv_rhs => rw [← Nat.cast_one] + rw [← natCast_lt_toENat_iff] + simp only [ENat.card, Nat.cast_one] + +end ENat + + namespace PartENat /-- `PartENat.card α` is the cardinality of `α` as an extended natural number. If `α` is infinite, `PartENat.card α = ⊤`. -/ +@[deprecated ENat.card (since := "2024-11-30")] def card (α : Type*) : PartENat := toPartENat (mk α) -@[simp] +-- The remainder of this section is about the deprecated `PartENat.card`. +set_option linter.deprecated false + +@[simp, deprecated ENat.card_eq_coe_fintype_card (since := "2024-11-30")] theorem card_eq_coe_fintype_card [Fintype α] : card α = Fintype.card α := mk_toPartENat_eq_coe_card -@[simp] +@[simp, deprecated ENat.card_eq_top_of_infinite (since := "2024-11-30")] theorem card_eq_top_of_infinite [Infinite α] : card α = ⊤ := mk_toPartENat_of_infinite -@[simp] +@[simp, deprecated ENat.card_sum (since := "2024-11-30")] theorem card_sum (α β : Type*) : PartENat.card (α ⊕ β) = PartENat.card α + PartENat.card β := by simp only [PartENat.card, Cardinal.mk_sum, map_add, Cardinal.toPartENat_lift] +@[deprecated ENat.card_congr (since := "2024-11-30")] theorem card_congr {α : Type*} {β : Type*} (f : α ≃ β) : PartENat.card α = PartENat.card β := Cardinal.toPartENat_congr f -@[simp] lemma card_ulift (α : Type*) : card (ULift α) = card α := card_congr Equiv.ulift +@[simp, deprecated ENat.card_ulift (since := "2024-11-30")] +lemma card_ulift (α : Type*) : card (ULift α) = card α := card_congr Equiv.ulift -@[simp] lemma card_plift (α : Type*) : card (PLift α) = card α := card_congr Equiv.plift +@[simp, deprecated ENat.card_plift (since := "2024-11-30")] +lemma card_plift (α : Type*) : card (PLift α) = card α := card_congr Equiv.plift +@[deprecated ENat.card_image_of_injOn (since := "2024-11-30")] theorem card_image_of_injOn {α : Type u} {β : Type v} {f : α → β} {s : Set α} (h : Set.InjOn f s) : card (f '' s) = card s := card_congr (Equiv.Set.imageOfInjOn f s h).symm +@[deprecated ENat.card_image_of_injective (since := "2024-11-30")] theorem card_image_of_injective {α : Type u} {β : Type v} (f : α → β) (s : Set α) (h : Function.Injective f) : card (f '' s) = card s := card_image_of_injOn h.injOn -- Should I keep the 6 following lemmas ? -- TODO: Add ofNat, zero, and one versions for simp confluence -@[simp] +@[simp, deprecated Cardinal.natCast_le_toENat_iff (since := "2024-11-30")] theorem _root_.Cardinal.natCast_le_toPartENat_iff {n : ℕ} {c : Cardinal} : ↑n ≤ toPartENat c ↔ ↑n ≤ c := by rw [← toPartENat_natCast n, toPartENat_le_iff_of_le_aleph0 (le_of_lt (nat_lt_aleph0 n))] -@[simp] +@[simp, deprecated Cardinal.toENat_le_natCast_iff (since := "2024-11-30")] theorem _root_.Cardinal.toPartENat_le_natCast_iff {c : Cardinal} {n : ℕ} : toPartENat c ≤ n ↔ c ≤ n := by rw [← toPartENat_natCast n, toPartENat_le_iff_of_lt_aleph0 (nat_lt_aleph0 n)] -@[simp] +@[simp, deprecated Cardinal.natCast_eq_toENat_iff (since := "2024-11-30")] theorem _root_.Cardinal.natCast_eq_toPartENat_iff {n : ℕ} {c : Cardinal} : ↑n = toPartENat c ↔ ↑n = c := by rw [le_antisymm_iff, le_antisymm_iff, Cardinal.toPartENat_le_natCast_iff, Cardinal.natCast_le_toPartENat_iff] -@[simp] +@[simp, deprecated Cardinal.toENat_eq_natCast_iff (since := "2024-11-30")] theorem _root_.Cardinal.toPartENat_eq_natCast_iff {c : Cardinal} {n : ℕ} : Cardinal.toPartENat c = n ↔ c = n := by rw [eq_comm, Cardinal.natCast_eq_toPartENat_iff, eq_comm] -@[simp] +@[simp, deprecated Cardinal.natCast_lt_toENat_iff (since := "2024-11-30")] theorem _root_.Cardinal.natCast_lt_toPartENat_iff {n : ℕ} {c : Cardinal} : ↑n < toPartENat c ↔ ↑n < c := by simp only [← not_le, Cardinal.toPartENat_le_natCast_iff] -@[simp] +@[simp, deprecated Cardinal.toENat_lt_natCast_iff (since := "2024-11-30")] theorem _root_.Cardinal.toPartENat_lt_natCast_iff {n : ℕ} {c : Cardinal} : toPartENat c < ↑n ↔ c < ↑n := by simp only [← not_le, Cardinal.natCast_le_toPartENat_iff] +@[deprecated ENat.card_eq_zero_iff_empty (since := "2024-11-30")] theorem card_eq_zero_iff_empty (α : Type*) : card α = 0 ↔ IsEmpty α := by rw [← Cardinal.mk_eq_zero_iff] conv_rhs => rw [← Nat.cast_zero] simp only [← Cardinal.toPartENat_eq_natCast_iff] simp only [PartENat.card, Nat.cast_zero] +@[deprecated ENat.card_le_one_iff_subsingleton (since := "2024-11-30")] theorem card_le_one_iff_subsingleton (α : Type*) : card α ≤ 1 ↔ Subsingleton α := by rw [← le_one_iff_subsingleton] conv_rhs => rw [← Nat.cast_one] rw [← Cardinal.toPartENat_le_natCast_iff] simp only [PartENat.card, Nat.cast_one] +@[deprecated ENat.one_lt_card_iff_nontrivial (since := "2024-11-30")] theorem one_lt_card_iff_nontrivial (α : Type*) : 1 < card α ↔ Nontrivial α := by rw [← Cardinal.one_lt_iff_nontrivial] conv_rhs => rw [← Nat.cast_one] From c769c4fbf51cfd9b36252153141ff0465df14bc2 Mon Sep 17 00:00:00 2001 From: leanprover-community-bot-assistant Date: Sun, 1 Dec 2024 00:37:03 +0000 Subject: [PATCH 438/829] chore(scripts): update nolints.json (#19647) I am happy to remove some nolints for you! --- scripts/nolints.json | 1 - 1 file changed, 1 deletion(-) diff --git a/scripts/nolints.json b/scripts/nolints.json index a698a174e1354..4f5d00a571a78 100644 --- a/scripts/nolints.json +++ b/scripts/nolints.json @@ -61,7 +61,6 @@ ["docBlame", "AffineMap.toFun"], ["docBlame", "AlgHom.toAddMonoidHom'"], ["docBlame", "AlgHom.toMonoidHom'"], - ["docBlame", "AlgebraCat.carrier"], ["docBlame", "Bifunctor.bimap"], ["docBlame", "Bimod.X"], ["docBlame", "Bimod.actLeft"], From a1f01bd5d16151d917c810909a2aa2e7b7ad64a0 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 1 Dec 2024 01:39:51 +0000 Subject: [PATCH 439/829] refactor(Algebra/Bilinear): generalize to non-commutative base rings (#19232) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This removes unnecessary `IsScalarTower R A A` and `SMulCommClass R A A` arguments from `LinearMap.mulLeft` and `LinearMap.mulRight`, respectively. As described in the new docstring, this allows respective specializations to `R := Aᵐᵒᵖ` and `R := A`. Unfortunately all of the lemmas have to be reordered to take advantage of these relaxed assumptions. In the chaos, a few lemmas also became generalized to non-unital rings. There is some overlap here with `DistribMulAction.toLinearMap`, but that doesn't work in the case of non-unital rings. --- Mathlib/Algebra/Algebra/Bilinear.lean | 251 +++++++++++++----------- Mathlib/RingTheory/PiTensorProduct.lean | 5 + 2 files changed, 146 insertions(+), 110 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Bilinear.lean b/Mathlib/Algebra/Algebra/Bilinear.lean index 02c54d00368ef..2abf893a86632 100644 --- a/Mathlib/Algebra/Algebra/Bilinear.lean +++ b/Mathlib/Algebra/Algebra/Bilinear.lean @@ -21,13 +21,71 @@ open TensorProduct Module namespace LinearMap section NonUnitalNonAssoc +variable (R A : Type*) -variable (R A : Type*) [CommSemiring R] [NonUnitalNonAssocSemiring A] [Module R A] - [SMulCommClass R A A] [IsScalarTower R A A] +section one_side +variable [Semiring R] [NonUnitalNonAssocSemiring A] [Module R A] + +section left +variable {A} [SMulCommClass R A A] + +/-- The multiplication on the left in a algebra is a linear map. + +Note that this only assumes `SMulCommClass R A A`, so that it also works for `R := Aᵐᵒᵖ`. + +When `A` is unital and associative, this is the same as `DistribMulAction.toLinearMap R A a` -/ +def mulLeft (a : A) : A →ₗ[R] A where + toFun := (a * ·) + map_add' := mul_add _ + map_smul' _ := mul_smul_comm _ _ + +@[simp] +theorem mulLeft_apply (a b : A) : mulLeft R a b = a * b := rfl + +@[simp] +theorem mulLeft_toAddMonoidHom (a : A) : (mulLeft R a : A →+ A) = AddMonoidHom.mulLeft a := rfl + +variable (A) in +@[simp] +theorem mulLeft_zero_eq_zero : mulLeft R (0 : A) = 0 := ext fun _ => zero_mul _ + +end left + +section right +variable {A} [IsScalarTower R A A] + +/-- The multiplication on the right in an algebra is a linear map. + +Note that this only assumes `IsScalarTower R A A`, so that it also works for `R := A`. + +When `A` is unital and associative, this is the same as +`DistribMulAction.toLinearMap R A (MulOpposite.op b)`. -/ +def mulRight (b : A) : A →ₗ[R] A where + toFun := (· * b) + map_add' _ _ := add_mul _ _ _ + map_smul' _ _ := smul_mul_assoc _ _ _ + +@[simp] +theorem mulRight_apply (a b : A) : mulRight R a b = b * a := rfl + +@[simp] +theorem mulRight_toAddMonoidHom (a : A) : (mulRight R a : A →+ A) = AddMonoidHom.mulRight a := rfl + +variable (A) in +@[simp] +theorem mulRight_zero_eq_zero : mulRight R (0 : A) = 0 := ext fun _ => mul_zero _ + +end right + +end one_side + +variable [CommSemiring R] [NonUnitalNonAssocSemiring A] [Module R A] +variable [SMulCommClass R A A] [IsScalarTower R A A] /-- The multiplication in a non-unital non-associative algebra is a bilinear map. A weaker version of this for semirings exists as `AddMonoidHom.mul`. -/ +@[simps!] def mul : A →ₗ[R] A →ₗ[R] A := LinearMap.mk₂ R (· * ·) add_mul smul_mul_assoc mul_add mul_smul_comm @@ -38,40 +96,16 @@ noncomputable def mul' : A ⊗[R] A →ₗ[R] A := variable {A} -/-- The multiplication on the left in a non-unital algebra is a linear map. -/ -def mulLeft (a : A) : A →ₗ[R] A := - mul R A a - -/-- The multiplication on the right in an algebra is a linear map. -/ -def mulRight (a : A) : A →ₗ[R] A := - (mul R A).flip a - /-- Simultaneous multiplication on the left and right is a linear map. -/ def mulLeftRight (ab : A × A) : A →ₗ[R] A := (mulRight R ab.snd).comp (mulLeft R ab.fst) -@[simp] -theorem mulLeft_toAddMonoidHom (a : A) : (mulLeft R a : A →+ A) = AddMonoidHom.mulLeft a := - rfl - -@[simp] -theorem mulRight_toAddMonoidHom (a : A) : (mulRight R a : A →+ A) = AddMonoidHom.mulRight a := - rfl - variable {R} @[simp] theorem mul_apply' (a b : A) : mul R A a b = a * b := rfl -@[simp] -theorem mulLeft_apply (a b : A) : mulLeft R a b = a * b := - rfl - -@[simp] -theorem mulRight_apply (a b : A) : mulRight R a b = b * a := - rfl - @[simp] theorem mulLeftRight_apply (a b x : A) : mulLeftRight R (a, b) x = a * x * b := rfl @@ -80,35 +114,41 @@ theorem mulLeftRight_apply (a b x : A) : mulLeftRight R (a, b) x = a * x * b := theorem mul'_apply {a b : A} : mul' R A (a ⊗ₜ b) = a * b := rfl -@[simp] -theorem mulLeft_zero_eq_zero : mulLeft R (0 : A) = 0 := - (mul R A).map_zero +end NonUnitalNonAssoc + +section NonUnital +variable (R A B : Type*) + +section one_side +variable [Semiring R] [NonUnitalSemiring A] [NonUnitalSemiring B] [Module R B] [Module R A] @[simp] -theorem mulRight_zero_eq_zero : mulRight R (0 : A) = 0 := - (mul R A).flip.map_zero +theorem mulLeft_mul [SMulCommClass R A A] (a b : A) : + mulLeft R (a * b) = (mulLeft R a).comp (mulLeft R b) := by + ext + simp only [mulLeft_apply, comp_apply, mul_assoc] -end NonUnitalNonAssoc +@[simp] +theorem mulRight_mul [IsScalarTower R A A] (a b : A) : + mulRight R (a * b) = (mulRight R b).comp (mulRight R a) := by + ext + simp only [mulRight_apply, comp_apply, mul_assoc] -section NonUnital +end one_side -variable (R A : Type*) [CommSemiring R] [NonUnitalSemiring A] [Module R A] [SMulCommClass R A A] - [IsScalarTower R A A] +variable [CommSemiring R] [NonUnitalSemiring A] [NonUnitalSemiring B] [Module R B] [Module R A] +variable [SMulCommClass R A A] [IsScalarTower R A A] +variable [SMulCommClass R B B] [IsScalarTower R B B] /-- The multiplication in a non-unital algebra is a bilinear map. A weaker version of this for non-unital non-associative algebras exists as `LinearMap.mul`. -/ -def _root_.NonUnitalAlgHom.lmul : A →ₙₐ[R] End R A := - { mul R A with - map_mul' := by - intro a b - ext c - exact mul_assoc a b c - map_zero' := by - ext a - exact zero_mul a } +def _root_.NonUnitalAlgHom.lmul : A →ₙₐ[R] End R A where + __ := mul R A + map_mul' := mulLeft_mul _ _ + map_zero' := mulLeft_zero_eq_zero _ _ -variable {R A} +variable {R A B} @[simp] theorem _root_.NonUnitalAlgHom.coe_lmul_eq_mul : ⇑(NonUnitalAlgHom.lmul R A) = mul R A := @@ -118,23 +158,6 @@ theorem commute_mulLeft_right (a b : A) : Commute (mulLeft R a) (mulRight R b) : ext c exact (mul_assoc a c b).symm -@[simp] -theorem mulLeft_mul (a b : A) : mulLeft R (a * b) = (mulLeft R a).comp (mulLeft R b) := by - ext - simp only [mulLeft_apply, comp_apply, mul_assoc] - -@[simp] -theorem mulRight_mul (a b : A) : mulRight R (a * b) = (mulRight R b).comp (mulRight R a) := by - ext - simp only [mulRight_apply, comp_apply, mul_assoc] - -end NonUnital - -section Semiring - -variable (R A B : Type*) [CommSemiring R] [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] - -variable {R A B} in /-- A `LinearMap` preserves multiplication if pre- and post- composition with `LinearMap.mul` are equivalent. By converting the statement into an equality of `LinearMap`s, this lemma allows various specialized `ext` lemmas about `→ₗ[R]` to then be applied. @@ -145,76 +168,84 @@ theorem map_mul_iff (f : A →ₗ[R] B) : (LinearMap.mul R A).compr₂ f = (LinearMap.mul R B ∘ₗ f).compl₂ f := Iff.symm LinearMap.ext_iff₂ -/-- The multiplication in an algebra is an algebra homomorphism into the endomorphisms on -the algebra. - -A weaker version of this for non-unital algebras exists as `NonUnitalAlgHom.mul`. -/ -def _root_.Algebra.lmul : A →ₐ[R] End R A := - { LinearMap.mul R A with - map_one' := by - ext a - exact one_mul a - map_mul' := by - intro a b - ext c - exact mul_assoc a b c - map_zero' := by - ext a - exact zero_mul a - commutes' := by - intro r - ext a - exact (Algebra.smul_def r a).symm } +end NonUnital -variable {R A} +section Semiring -@[simp] -theorem _root_.Algebra.coe_lmul_eq_mul : ⇑(Algebra.lmul R A) = mul R A := - rfl +variable (R A : Type*) +section one_side +variable [Semiring R] [Semiring A] -theorem _root_.Algebra.lmul_injective : Function.Injective (Algebra.lmul R A) := - fun a₁ a₂ h ↦ by simpa using DFunLike.congr_fun h 1 +section left +variable [Module R A] [SMulCommClass R A A] -theorem _root_.Algebra.lmul_isUnit_iff {x : A} : - IsUnit (Algebra.lmul R A x) ↔ IsUnit x := by - rw [Module.End_isUnit_iff, Iff.comm] - exact IsUnit.isUnit_iff_mulLeft_bijective +@[simp] +theorem mulLeft_one : mulLeft R (1 : A) = LinearMap.id := ext fun _ => one_mul _ @[simp] theorem mulLeft_eq_zero_iff (a : A) : mulLeft R a = 0 ↔ a = 0 := by constructor <;> intro h -- Porting note: had to supply `R` explicitly in `@mulLeft_apply` below - · rw [← mul_one a, ← @mulLeft_apply R _ _ _ _ _ _ a 1, h, LinearMap.zero_apply] + · rw [← mul_one a, ← mulLeft_apply R a 1, h, LinearMap.zero_apply] · rw [h] - exact mulLeft_zero_eq_zero + exact mulLeft_zero_eq_zero _ _ + +@[simp] +theorem pow_mulLeft (a : A) (n : ℕ) : mulLeft R a ^ n = mulLeft R (a ^ n) := + match n with + | 0 => by rw [pow_zero, pow_zero, mulLeft_one, LinearMap.one_eq_id] + | (n + 1) => by rw [pow_succ, pow_succ, mulLeft_mul, LinearMap.mul_eq_comp, pow_mulLeft] + +end left + +section right +variable [Module R A] [IsScalarTower R A A] + +@[simp] +theorem mulRight_one : mulRight R (1 : A) = LinearMap.id := ext fun _ => mul_one _ @[simp] theorem mulRight_eq_zero_iff (a : A) : mulRight R a = 0 ↔ a = 0 := by constructor <;> intro h -- Porting note: had to supply `R` explicitly in `@mulRight_apply` below - · rw [← one_mul a, ← @mulRight_apply R _ _ _ _ _ _ a 1, h, LinearMap.zero_apply] + · rw [← one_mul a, ← mulRight_apply R a 1, h, LinearMap.zero_apply] · rw [h] - exact mulRight_zero_eq_zero + exact mulRight_zero_eq_zero _ _ @[simp] -theorem mulLeft_one : mulLeft R (1 : A) = LinearMap.id := by - ext - simp +theorem pow_mulRight (a : A) (n : ℕ) : mulRight R a ^ n = mulRight R (a ^ n) := + match n with + | 0 => by rw [pow_zero, pow_zero, mulRight_one, LinearMap.one_eq_id] + | (n + 1) => by rw [pow_succ, pow_succ', mulRight_mul, LinearMap.mul_eq_comp, pow_mulRight] -@[simp] -theorem mulRight_one : mulRight R (1 : A) = LinearMap.id := by - ext - simp +end right -@[simp] -theorem pow_mulLeft (a : A) (n : ℕ) : mulLeft R a ^ n = mulLeft R (a ^ n) := by - simpa only [mulLeft, ← Algebra.coe_lmul_eq_mul] using (map_pow (Algebra.lmul R A) a n).symm +end one_side + +variable [CommSemiring R] [Semiring A] [Algebra R A] + +/-- The multiplication in an algebra is an algebra homomorphism into the endomorphisms on +the algebra. + +A weaker version of this for non-unital algebras exists as `NonUnitalAlgHom.lmul`. -/ +def _root_.Algebra.lmul : A →ₐ[R] End R A where + __ := NonUnitalAlgHom.lmul R A + map_one' := mulLeft_one _ _ + commutes' r := ext fun a => (Algebra.smul_def r a).symm + +variable {R A} @[simp] -theorem pow_mulRight (a : A) (n : ℕ) : mulRight R a ^ n = mulRight R (a ^ n) := by - simp only [mulRight, ← Algebra.coe_lmul_eq_mul] - exact - LinearMap.coe_injective (((mulRight R a).coe_pow n).symm ▸ mul_right_iterate a n) +theorem _root_.Algebra.coe_lmul_eq_mul : ⇑(Algebra.lmul R A) = mul R A := + rfl + +theorem _root_.Algebra.lmul_injective : Function.Injective (Algebra.lmul R A) := + fun a₁ a₂ h ↦ by simpa using DFunLike.congr_fun h 1 + +theorem _root_.Algebra.lmul_isUnit_iff {x : A} : + IsUnit (Algebra.lmul R A x) ↔ IsUnit x := by + rw [Module.End_isUnit_iff, Iff.comm] + exact IsUnit.isUnit_iff_mulLeft_bijective theorem toSpanSingleton_eq_algebra_linearMap : toSpanSingleton R A 1 = Algebra.linearMap R A := by ext; simp diff --git a/Mathlib/RingTheory/PiTensorProduct.lean b/Mathlib/RingTheory/PiTensorProduct.lean index 59dfb88d4b256..26d7db3188d1e 100644 --- a/Mathlib/RingTheory/PiTensorProduct.lean +++ b/Mathlib/RingTheory/PiTensorProduct.lean @@ -272,6 +272,11 @@ noncomputable def constantBaseRingEquiv : (⨂[R] _ : ι, R) ≃ₐ[R] R := toFun ((lift.tprod _).trans Finset.prod_const_one) (by + -- one of these is required, the other is a performance optimization + letI : IsScalarTower R (⨂[R] x : ι, R) (⨂[R] x : ι, R) := + IsScalarTower.right (R := R) (A := ⨂[R] (x : ι), R) + letI : SMulCommClass R (⨂[R] x : ι, R) (⨂[R] x : ι, R) := + Algebra.to_smulCommClass (R := R) (A := ⨂[R] x : ι, R) rw [LinearMap.map_mul_iff] ext x y show toFun (tprod R x * tprod R y) = toFun (tprod R x) * toFun (tprod R y) From b8b8a288609c214ad2aac7d3837bedce7e2b4875 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sun, 1 Dec 2024 08:03:37 +0000 Subject: [PATCH 440/829] chore: update benchmark runner for leanprover/lean4#5684 (#19649) leanprover/lean4#5684 (which landed in v4.14.0-rc3) has removed the `--lean` flag for `lake`, indicating that we should use the `LEAN` environment variable instead. --- scripts/bench/temci-config.run.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/scripts/bench/temci-config.run.yml b/scripts/bench/temci-config.run.yml index 679eb6a0e6d8f..49179c7e7dad2 100644 --- a/scripts/bench/temci-config.run.yml +++ b/scripts/bench/temci-config.run.yml @@ -7,7 +7,7 @@ rusage_properties: ['maxrss'] cmd: | # use build cache for proofwidgets, but not for anything else - bash -c 'set -eo pipefail; lake clean 1>&2 && LEAN_PATH=$(lean --print-libdir) lake build proofwidgets 1>&2 && rm -f .lake/packages/batteries/.lake/build/bin/runLinter 1>&2 && lake build --no-cache -v --lean ./scripts/bench/fake-root/bin/lean | ./scripts/bench/accumulate_profile.py | grep -v took' + bash -c 'set -eo pipefail; lake clean 1>&2 && LEAN_PATH=$(lean --print-libdir) lake build proofwidgets 1>&2 && rm -f .lake/packages/batteries/.lake/build/bin/runLinter 1>&2 && LEAN=./scripts/bench/fake-root/bin/lean lake build --no-cache -v | ./scripts/bench/accumulate_profile.py | grep -v took' parse_output: true runs: 1 - attributes: From 81b5c5ca3e243ed6392cb406cdc0fd6449ff0c89 Mon Sep 17 00:00:00 2001 From: Thmoas-Guan <150537269+Thmoas-Guan@users.noreply.github.com> Date: Sun, 1 Dec 2024 09:31:07 +0000 Subject: [PATCH 441/829] feat(FieldTheory/Galois): Lemmas of galois theory (#16979) Add lemmas about Galois extensions, especially the special case of the Galois correspondence. Co-authored-by: Yongle Hu Nailin Guan <3571819596@qq.com> Jingting Wang --- Mathlib/Algebra/Group/Subgroup/Pointwise.lean | 12 +++- Mathlib/FieldTheory/Galois/Basic.lean | 68 ++++++++++++++++++- .../FieldTheory/IntermediateField/Basic.lean | 13 ++++ Mathlib/FieldTheory/Normal.lean | 4 ++ 4 files changed, 94 insertions(+), 3 deletions(-) diff --git a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean index da9423cd8aba9..d5827d8bec2d1 100644 --- a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean +++ b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean @@ -411,8 +411,7 @@ theorem subgroup_mul_singleton {H : Subgroup G} {h : G} (hh : h ∈ H) : (H : Se theorem singleton_mul_subgroup {H : Subgroup G} {h : G} (hh : h ∈ H) : {h} * (H : Set G) = H := by simp [preimage, mul_mem_cancel_left (inv_mem hh)] -theorem Normal.conjAct {G : Type*} [Group G] {H : Subgroup G} (hH : H.Normal) (g : ConjAct G) : - g • H = H := +theorem Normal.conjAct {H : Subgroup G} (hH : H.Normal) (g : ConjAct G) : g • H = H := have : ∀ g : ConjAct G, g • H ≤ H := fun _ => map_le_iff_le_comap.2 fun _ h => hH.conj_mem _ h _ (this g).antisymm <| (smul_inv_smul g H).symm.trans_le (map_mono <| this _) @@ -421,6 +420,15 @@ 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 Normal.of_conjugate_fixed {H : Subgroup G} (h : ∀ g : G, (MulAut.conj g) • H = H) : + H.Normal := by + constructor + intro n hn g + rw [← h g, Subgroup.mem_pointwise_smul_iff_inv_smul_mem, ← map_inv, MulAut.smul_def, + MulAut.conj_apply, inv_inv, mul_assoc, mul_assoc, inv_mul_cancel, mul_one, + ← mul_assoc, inv_mul_cancel, one_mul] + exact hn + theorem normalCore_eq_iInf_conjAct (H : Subgroup G) : H.normalCore = ⨅ (g : ConjAct G), g • H := by ext g diff --git a/Mathlib/FieldTheory/Galois/Basic.lean b/Mathlib/FieldTheory/Galois/Basic.lean index 22fbc7307b2ad..86c87ea00d274 100644 --- a/Mathlib/FieldTheory/Galois/Basic.lean +++ b/Mathlib/FieldTheory/Galois/Basic.lean @@ -1,7 +1,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 +Authors: Thomas Browning, Patrick Lutz, Yongle Hu, Jingting Wang -/ import Mathlib.FieldTheory.Fixed import Mathlib.FieldTheory.NormalClosure @@ -289,6 +289,72 @@ def galoisCoinsertionIntermediateFieldSubgroup [FiniteDimensional F E] [IsGalois end IsGalois +section + +/-In this section we prove that the normal subgroups correspond to the Galois subextensions +in the Galois correspondence and its related results.-/ + +variable {K L : Type*} [Field K] [Field L] [Algebra K L] + +open IntermediateField + +open scoped Pointwise + +lemma IntermediateField.restrictNormalHom_ker (E : IntermediateField K L) [Normal K E] : + (restrictNormalHom E).ker = E.fixingSubgroup := by + simp [fixingSubgroup, Subgroup.ext_iff, AlgEquiv.ext_iff, Subtype.ext_iff, + restrictNormalHom_apply, mem_fixingSubgroup_iff] + +namespace IsGalois + +variable (E : IntermediateField K L) + +/-- If `H` is a normal Subgroup of `Gal(L / K)`, then `fixedField H` is Galois over `K`. -/ +instance of_fixedField_normal_subgroup [IsGalois K L] + (H : Subgroup (L ≃ₐ[K] L)) [hn : Subgroup.Normal H] : IsGalois K (fixedField H) where + to_isSeparable := Algebra.isSeparable_tower_bot_of_isSeparable K (fixedField H) L + to_normal := by + apply normal_iff_forall_map_le'.mpr + rintro σ x ⟨a, ha, rfl⟩ τ + exact (symm_apply_eq σ).mp (ha ⟨σ⁻¹ * τ * σ, Subgroup.Normal.conj_mem' hn τ.1 τ.2 σ⟩) + +/-- If `H` is a normal Subgroup of `Gal(L / K)`, then `Gal(fixedField H / K)` is isomorphic to +`Gal(L / K) ⧸ H`. -/ +noncomputable def normalAutEquivQuotient [FiniteDimensional K L] [IsGalois K L] + (H : Subgroup (L ≃ₐ[K] L)) [Subgroup.Normal H] : + (L ≃ₐ[K] L) ⧸ H ≃* ((fixedField H) ≃ₐ[K] (fixedField H)) := + (QuotientGroup.quotientMulEquivOfEq ((fixingSubgroup_fixedField H).symm.trans + (fixedField H).restrictNormalHom_ker.symm)).trans <| + QuotientGroup.quotientKerEquivOfSurjective (restrictNormalHom (fixedField H)) <| + restrictNormalHom_surjective L + +lemma normalAutEquivQuotient_apply [FiniteDimensional K L] [IsGalois K L] + (H : Subgroup (L ≃ₐ[K] L)) [Subgroup.Normal H] (σ : (L ≃ₐ[K] L)) : + normalAutEquivQuotient H σ = (restrictNormalHom (fixedField H)) σ := rfl + +open scoped Pointwise + +@[simp] +theorem map_fixingSubgroup (σ : L ≃ₐ[K] L) : + (E.map σ).fixingSubgroup = (MulAut.conj σ) • E.fixingSubgroup := by + ext τ + simp only [coe_map, AlgHom.coe_coe, Set.mem_image, SetLike.mem_coe, AlgEquiv.smul_def, + forall_exists_index, and_imp, forall_apply_eq_imp_iff₂, Subtype.forall, + Subgroup.mem_pointwise_smul_iff_inv_smul_mem, ← symm_apply_eq, + IntermediateField.fixingSubgroup, mem_fixingSubgroup_iff] + rfl + +/-- Let `E` be an intermediateField of a Galois extension `L / K`. If `E / K` is +Galois extension, then `E.fixingSubgroup` is a normal subgroup of `Gal(L / K)`. -/ +instance fixingSubgroup_normal_of_isGalois [IsGalois K L] [IsGalois K E] : + E.fixingSubgroup.Normal := by + apply Subgroup.Normal.of_conjugate_fixed (fun σ ↦ ?_) + rw [← map_fixingSubgroup, normal_iff_forall_map_eq'.mp inferInstance σ] + +end IsGalois + +end + end GaloisCorrespondence section GaloisEquivalentDefinitions diff --git a/Mathlib/FieldTheory/IntermediateField/Basic.lean b/Mathlib/FieldTheory/IntermediateField/Basic.lean index 33d4aa8b3bf97..9a52885edff92 100644 --- a/Mathlib/FieldTheory/IntermediateField/Basic.lean +++ b/Mathlib/FieldTheory/IntermediateField/Basic.lean @@ -613,6 +613,19 @@ theorem mem_lift {F : IntermediateField K L} {E : IntermediateField K F} (x : F) x.1 ∈ lift E ↔ x ∈ E := Subtype.val_injective.mem_set_image +/--The algEquiv between an intermediate field and its lift-/ +def liftAlgEquiv {E : IntermediateField K L} (F : IntermediateField K E) : ↥F ≃ₐ[K] lift F where + toFun x := ⟨x.1.1, (mem_lift x.1).mpr x.2⟩ + invFun x := ⟨⟨x.1, lift_le F x.2⟩, (mem_lift ⟨x.1, lift_le F x.2⟩).mp x.2⟩ + left_inv := congrFun rfl + right_inv := congrFun rfl + map_mul' _ _ := rfl + map_add' _ _ := rfl + commutes' _ := rfl + +lemma liftAlgEquiv_apply {E : IntermediateField K L} (F : IntermediateField K E) (x : F) : + (liftAlgEquiv F x).1 = x := rfl + section RestrictScalars variable (K) diff --git a/Mathlib/FieldTheory/Normal.lean b/Mathlib/FieldTheory/Normal.lean index dc0da73493444..12486d2a18034 100644 --- a/Mathlib/FieldTheory/Normal.lean +++ b/Mathlib/FieldTheory/Normal.lean @@ -290,6 +290,10 @@ theorem AlgEquiv.restrictNormal_trans [Normal F E] : def AlgEquiv.restrictNormalHom [Normal F E] : (K₁ ≃ₐ[F] K₁) →* E ≃ₐ[F] E := MonoidHom.mk' (fun χ => χ.restrictNormal E) fun ω χ => χ.restrictNormal_trans ω E +lemma AlgEquiv.restrictNormalHom_apply (L : IntermediateField F K₁) [Normal F L] + (σ : (K₁ ≃ₐ[F] K₁)) (x : L) : restrictNormalHom L σ x = σ x := + AlgEquiv.restrictNormal_commutes σ L x + variable (F K₁) /-- If `K₁/E/F` is a tower of fields with `E/F` normal then `AlgHom.restrictNormal'` is an From c820366725967a74108c9b211ec923469e94bd5b Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Sun, 1 Dec 2024 09:59:41 +0000 Subject: [PATCH 442/829] feat(Algebra/MvPolynomial/Degrees): degreeOf_mul_X_eq_degreeOf_add_one_iff (#17553) Co-authored-by: Moritz Firsching --- Mathlib/Algebra/MvPolynomial/Degrees.lean | 14 +++++++++++++- Mathlib/Data/Finset/Lattice/Fold.lean | 5 +++++ 2 files changed, 18 insertions(+), 1 deletion(-) diff --git a/Mathlib/Algebra/MvPolynomial/Degrees.lean b/Mathlib/Algebra/MvPolynomial/Degrees.lean index 9dd0aa38991a3..e0e9da35be358 100644 --- a/Mathlib/Algebra/MvPolynomial/Degrees.lean +++ b/Mathlib/Algebra/MvPolynomial/Degrees.lean @@ -291,7 +291,6 @@ theorem degreeOf_mul_X_ne {i j : σ} (f : MvPolynomial σ R) (h : i ≠ j) : simp only [Finsupp.single, Nat.one_ne_zero, add_right_eq_self, addRightEmbedding_apply, coe_mk, Pi.add_apply, comp_apply, ite_eq_right_iff, Finsupp.coe_add, Pi.single_eq_of_ne h] --- TODO in the following we have equality iff `f ≠ 0` theorem degreeOf_mul_X_eq (j : σ) (f : MvPolynomial σ R) : degreeOf j (f * X j) ≤ degreeOf j f + 1 := by classical @@ -301,6 +300,19 @@ theorem degreeOf_mul_X_eq (j : σ) (f : MvPolynomial σ R) : convert Multiset.count_le_of_le j (degrees_X' (R := R) j) rw [Multiset.count_singleton_self] +theorem degreeOf_mul_X_eq_degreeOf_add_one_iff (j : σ) (f : MvPolynomial σ R) : + degreeOf j (f * X j) = degreeOf j f + 1 ↔ f ≠ 0 := by + refine ⟨fun h => by by_contra ha; simp [ha] at h, fun h => ?_⟩ + apply Nat.le_antisymm (degreeOf_mul_X_eq j f) + have : (f.support.sup fun m ↦ m j) + 1 = (f.support.sup fun m ↦ (m j + 1)) := + Finset.comp_sup_eq_sup_comp_of_nonempty @Nat.succ_le_succ (support_nonempty.mpr h) + simp only [degreeOf_eq_sup, support_mul_X, this] + apply Finset.sup_le + intro x hx + simp only [Finset.sup_map, bot_eq_zero', add_pos_iff, zero_lt_one, or_true, Finset.le_sup_iff] + use x + simpa using mem_support_iff.mp hx + theorem degreeOf_C_mul_le (p : MvPolynomial σ R) (i : σ) (c : R) : (C c * p).degreeOf i ≤ p.degreeOf i := by unfold degreeOf diff --git a/Mathlib/Data/Finset/Lattice/Fold.lean b/Mathlib/Data/Finset/Lattice/Fold.lean index 404b2ac9105b3..a9cc52080c652 100644 --- a/Mathlib/Data/Finset/Lattice/Fold.lean +++ b/Mathlib/Data/Finset/Lattice/Fold.lean @@ -1106,6 +1106,11 @@ section LinearOrder variable [LinearOrder α] {s : Finset ι} (H : s.Nonempty) {f : ι → α} {a : α} +theorem comp_sup_eq_sup_comp_of_nonempty [OrderBot α] [SemilatticeSup β] [OrderBot β] + {g : α → β} (mono_g : Monotone g) (H : s.Nonempty) : g (s.sup f) = s.sup (g ∘ f) := by + rw [← Finset.sup'_eq_sup H, ← Finset.sup'_eq_sup H] + exact Finset.comp_sup'_eq_sup'_comp H g (fun x y ↦ Monotone.map_sup mono_g x y) + @[simp] theorem le_sup'_iff : a ≤ s.sup' H f ↔ ∃ b ∈ s, a ≤ f b := by rw [← WithBot.coe_le_coe, coe_sup', Finset.le_sup_iff (WithBot.bot_lt_coe a)] From 822e8564863729fa71322d896c9f6aa7903adef4 Mon Sep 17 00:00:00 2001 From: Xavier Roblot Date: Sun, 1 Dec 2024 09:59:43 +0000 Subject: [PATCH 443/829] feat(NumberTheory): Abel summation (#19289) Proof of several versions of Abel summation formula. This PR is part of the proof of the Analytic Class Number Formula. Co-authored-by: Xavier Roblot <46200072+xroblot@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/Algebra/BigOperators/Module.lean | 7 + .../Integral/IntervalIntegral.lean | 5 + Mathlib/NumberTheory/AbelSummation.lean | 173 ++++++++++++++++++ 4 files changed, 186 insertions(+) create mode 100644 Mathlib/NumberTheory/AbelSummation.lean diff --git a/Mathlib.lean b/Mathlib.lean index e7bf4ac3e66d4..f120c15dfeb9e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3761,6 +3761,7 @@ import Mathlib.ModelTheory.Syntax import Mathlib.ModelTheory.Types import Mathlib.ModelTheory.Ultraproducts import Mathlib.NumberTheory.ADEInequality +import Mathlib.NumberTheory.AbelSummation import Mathlib.NumberTheory.ArithmeticFunction import Mathlib.NumberTheory.Basic import Mathlib.NumberTheory.Bernoulli diff --git a/Mathlib/Algebra/BigOperators/Module.lean b/Mathlib/Algebra/BigOperators/Module.lean index 2329df72948f4..37171b85f8437 100644 --- a/Mathlib/Algebra/BigOperators/Module.lean +++ b/Mathlib/Algebra/BigOperators/Module.lean @@ -41,6 +41,13 @@ theorem sum_Ico_by_parts (hmn : m < n) : simp_rw [this, sum_neg_distrib, sum_range_succ, smul_add] abel +theorem sum_Ioc_by_parts (hmn : m < n) : + ∑ i ∈ Ioc m n, f i • g i = + f n • G (n + 1) - f (m + 1) • G (m + 1) + - ∑ i ∈ Ioc m (n - 1), (f (i + 1) - f i) • G (i + 1) := by + simpa only [← Nat.Ico_succ_succ, Nat.succ_eq_add_one, Nat.sub_add_cancel (Nat.one_le_of_lt hmn), + add_tsub_cancel_right] using sum_Ico_by_parts f g (Nat.succ_lt_succ hmn) + variable (n) /-- **Summation by parts** for ranges -/ diff --git a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean index 6870c55e4e416..0a214921062a9 100644 --- a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean @@ -85,6 +85,11 @@ theorem intervalIntegrable_iff : IntervalIntegrable f μ a b ↔ IntegrableOn f theorem IntervalIntegrable.def' (h : IntervalIntegrable f μ a b) : IntegrableOn f (Ι a b) μ := intervalIntegrable_iff.mp h +theorem IntervalIntegrable.congr {g : ℝ → E} (hf : IntervalIntegrable f μ a b) + (h : f =ᵐ[μ.restrict (Ι a b)] g) : + IntervalIntegrable g μ a b := by + rwa [intervalIntegrable_iff, ← integrableOn_congr_fun_ae h, ← intervalIntegrable_iff] + theorem intervalIntegrable_iff_integrableOn_Ioc_of_le (hab : a ≤ b) : IntervalIntegrable f μ a b ↔ IntegrableOn f (Ioc a b) μ := by rw [intervalIntegrable_iff, uIoc_of_le hab] diff --git a/Mathlib/NumberTheory/AbelSummation.lean b/Mathlib/NumberTheory/AbelSummation.lean new file mode 100644 index 0000000000000..69dc93ce98099 --- /dev/null +++ b/Mathlib/NumberTheory/AbelSummation.lean @@ -0,0 +1,173 @@ +/- +Copyright (c) 2024 Xavier Roblot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Xavier Roblot +-/ +import Mathlib.MeasureTheory.Integral.FundThmCalculus + +/-! +# Abel's summation formula + +We prove several versions of Abel's summation formula. + +## Results + +* `sum_mul_eq_sub_sub_integral_mul`: general statement of the formula for a sum between two +(nonnegative) reals `a` and `b`. + +* `sum_mul_eq_sub_integral_mul`: a specialized version of `sum_mul_eq_sub_sub_integral_mul` for + the case `a = 0`. + + * `sum_mul_eq_sub_integral_mul`: a specialized version of `sum_mul_eq_sub_integral_mul` for + when the first coefficient of the sequence is `0`. This is useful for `ArithmeticFunction`. + +## References + +* + +-/ + +noncomputable section + +open Finset intervalIntegral MeasureTheory IntervalIntegrable + +variable {𝕜 : Type*} [RCLike 𝕜] (c : ℕ → 𝕜) {f : ℝ → 𝕜} {a b : ℝ} + +namespace abelSummationProof + +private theorem sumlocc (n : ℕ) : + ∀ᵐ t, t ∈ Set.Icc (n : ℝ) (n + 1) → ∑ k ∈ Icc 0 ⌊t⌋₊, c k = ∑ k ∈ Icc 0 n, c k := by + filter_upwards [Ico_ae_eq_Icc] with t h ht + rw [Nat.floor_eq_on_Ico _ _ (h.mpr ht)] + +private theorem integralmulsum (hf_diff : ∀ t ∈ Set.Icc a b, DifferentiableAt ℝ f t) + (hf_int : IntegrableOn (deriv f) (Set.Icc a b)) (t₁ t₂ : ℝ) (n : ℕ) (h : t₁ ≤ t₂) + (h₁ : n ≤ t₁) (h₂ : t₂ ≤ n + 1) (h₃ : a ≤ t₁) (h₄ : t₂ ≤ b) : + ∫ t in t₁..t₂, deriv f t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k = + (f t₂ - f t₁) * ∑ k ∈ Icc 0 n, c k := by + have h_inc₁ : Ι t₁ t₂ ⊆ Set.Icc n (n + 1) := + Set.uIoc_of_le h ▸ Set.Ioc_subset_Icc_self.trans <| Set.Icc_subset_Icc h₁ h₂ + have h_inc₂ : Set.uIcc t₁ t₂ ⊆ Set.Icc a b := Set.uIcc_of_le h ▸ Set.Icc_subset_Icc h₃ h₄ + rw [← integral_deriv_eq_sub (fun t ht ↦ hf_diff t (h_inc₂ ht)), ← integral_mul_const] + · refine integral_congr_ae ?_ + filter_upwards [sumlocc c n] with t h h' + rw [h (h_inc₁ h')] + · refine (intervalIntegrable_iff_integrableOn_Icc_of_le h).mpr (hf_int.mono_set ?_) + rwa [← Set.uIcc_of_le h] + +private theorem ineqofmemIco {k : ℕ} (hk : k ∈ Set.Ico (⌊a⌋₊ + 1) ⌊b⌋₊) : + a ≤ k ∧ k + 1 ≤ b := by + constructor + · have := (Set.mem_Ico.mp hk).1 + exact le_of_lt <| (Nat.floor_lt' (by omega)).mp this + · rw [← Nat.cast_add_one, ← Nat.le_floor_iff' (Nat.succ_ne_zero k)] + exact (Set.mem_Ico.mp hk).2 + +private theorem ineqofmemIco' {k : ℕ} (hk : k ∈ Ico (⌊a⌋₊ + 1) ⌊b⌋₊) : + a ≤ k ∧ k + 1 ≤ b := + ineqofmemIco (by rwa [← Finset.coe_Ico]) + +private theorem integrablemulsum (ha : 0 ≤ a) (hb : ⌊a⌋₊ < ⌊b⌋₊) + (hf_int : IntegrableOn (deriv f) (Set.Icc a b)) : + IntegrableOn (fun t ↦ deriv f t * (∑ k ∈ Icc 0 ⌊t⌋₊, c k)) (Set.Icc a b) := by + have h_locint {t₁ t₂ : ℝ} {n : ℕ} (h : t₁ ≤ t₂) (h₁ : n ≤ t₁) (h₂ : t₂ ≤ n + 1) + (h₃ : a ≤ t₁) (h₄ : t₂ ≤ b) : + IntervalIntegrable (fun t ↦ deriv f t * (∑ k ∈ Icc 0 ⌊t⌋₊, c k)) volume t₁ t₂ := by + rw [intervalIntegrable_iff_integrableOn_Icc_of_le h] + exact (IntegrableOn.mono_set (hf_int.mul_const _) (Set.Icc_subset_Icc h₃ h₄)).congr + <| ae_restrict_of_ae_restrict_of_subset (Set.Icc_subset_Icc h₁ h₂) + <| (ae_restrict_iff' measurableSet_Icc).mpr + (by filter_upwards [sumlocc c n] with t h ht using by rw [h ht]) + have aux1 : 0 ≤ b := (Nat.pos_of_floor_pos <| (Nat.zero_le _).trans_lt hb).le + have aux2 : ⌊a⌋₊ + 1 ≤ b := by rwa [← Nat.cast_add_one, ← Nat.le_floor_iff aux1] + have aux3 : a ≤ ⌊a⌋₊ + 1 := (Nat.lt_floor_add_one _).le + have aux4 : a ≤ ⌊b⌋₊ := le_of_lt (by rwa [← Nat.floor_lt ha]) + -- now break up into 3 subintervals + rw [← intervalIntegrable_iff_integrableOn_Icc_of_le (aux3.trans aux2)] + have I1 : IntervalIntegrable _ volume a ↑(⌊a⌋₊ + 1) := + h_locint (mod_cast aux3) (Nat.floor_le ha) (mod_cast le_rfl) le_rfl (mod_cast aux2) + have I2 : IntervalIntegrable _ volume ↑(⌊a⌋₊ + 1) ⌊b⌋₊ := + trans_iterate_Ico hb fun k hk ↦ h_locint (mod_cast k.le_succ) + le_rfl (mod_cast le_rfl) (ineqofmemIco hk).1 (mod_cast (ineqofmemIco hk).2) + have I3 : IntervalIntegrable _ volume ⌊b⌋₊ b := + h_locint (Nat.floor_le aux1) le_rfl (Nat.lt_floor_add_one _).le aux4 le_rfl + exact (I1.trans I2).trans I3 + +/-- Abel's summation formula. -/ +theorem _root_.sum_mul_eq_sub_sub_integral_mul (ha : 0 ≤ a) (hab : a ≤ b) + (hf_diff : ∀ t ∈ Set.Icc a b, DifferentiableAt ℝ f t) + (hf_int : IntegrableOn (deriv f) (Set.Icc a b)) : + ∑ k ∈ Ioc ⌊a⌋₊ ⌊b⌋₊, f k * c k = + f b * (∑ k ∈ Icc 0 ⌊b⌋₊, c k) - f a * (∑ k ∈ Icc 0 ⌊a⌋₊, c k) - + ∫ t in Set.Ioc a b, deriv f t * (∑ k ∈ Icc 0 ⌊t⌋₊, c k) := by + rw [← integral_of_le hab] + have aux1 : ⌊a⌋₊ ≤ a := Nat.floor_le ha + have aux2 : b ≤ ⌊b⌋₊ + 1 := (Nat.lt_floor_add_one _).le + -- We consider two cases depending on whether the sum is empty or not + obtain hb | hb := eq_or_lt_of_le (Nat.floor_le_floor hab) + · rw [hb, Ioc_eq_empty_of_le le_rfl, sum_empty, ← sub_mul, + integralmulsum c hf_diff hf_int _ _ ⌊b⌋₊ hab (hb ▸ aux1) aux2 le_rfl le_rfl, sub_self] + have aux3 : a ≤ ⌊a⌋₊ + 1 := (Nat.lt_floor_add_one _).le + have aux4 : ⌊a⌋₊ + 1 ≤ b := by rwa [← Nat.cast_add_one, ← Nat.le_floor_iff (ha.trans hab)] + have aux5 : ⌊b⌋₊ ≤ b := Nat.floor_le (ha.trans hab) + have aux6 : a ≤ ⌊b⌋₊ := Nat.floor_lt ha |>.mp hb |>.le + simp_rw [← smul_eq_mul, sum_Ioc_by_parts (fun k ↦ f k) _ hb, range_eq_Ico, Nat.Ico_succ_right, + smul_eq_mul] + have : ∑ k ∈ Ioc ⌊a⌋₊ (⌊b⌋₊ - 1), (f ↑(k + 1) - f k) * ∑ n ∈ Icc 0 k, c n = + ∑ k ∈ Ico (⌊a⌋₊ + 1) ⌊b⌋₊, ∫ t in k..↑(k + 1), deriv f t * ∑ n ∈ Icc 0 ⌊t⌋₊, c n := by + rw [← Nat.Ico_succ_succ, Nat.succ_eq_add_one, Nat.succ_eq_add_one, Nat.sub_add_cancel + (by omega), Eq.comm] + exact sum_congr rfl fun k hk ↦ (integralmulsum c hf_diff hf_int _ _ _ (mod_cast k.le_succ) + le_rfl (mod_cast le_rfl) (ineqofmemIco' hk).1 <| mod_cast (ineqofmemIco' hk).2) + rw [this, sum_integral_adjacent_intervals_Ico hb, Nat.cast_add, Nat.cast_one, + ← integral_interval_sub_left (a := a) (c := ⌊a⌋₊ + 1), + ← integral_add_adjacent_intervals (b := ⌊b⌋₊) (c := b), + integralmulsum c hf_diff hf_int _ _ _ aux3 aux1 le_rfl le_rfl aux4, + integralmulsum c hf_diff hf_int _ _ _ aux5 le_rfl aux2 aux6 le_rfl] + · ring + -- now deal with the integrability side goals + -- (Note we have 5 goals, but the 1st and 3rd are identical. TODO: find a non-hacky way of dealing + -- with both at once.) + · rw [intervalIntegrable_iff_integrableOn_Icc_of_le aux6] + exact (integrablemulsum c ha hb hf_int).mono_set (Set.Icc_subset_Icc_right aux5) + · rw [intervalIntegrable_iff_integrableOn_Icc_of_le aux5] + exact (integrablemulsum c ha hb hf_int).mono_set (Set.Icc_subset_Icc_left aux6) + · rw [intervalIntegrable_iff_integrableOn_Icc_of_le aux6] + exact (integrablemulsum c ha hb hf_int).mono_set (Set.Icc_subset_Icc_right aux5) + · rw [intervalIntegrable_iff_integrableOn_Icc_of_le aux3] + exact (integrablemulsum c ha hb hf_int).mono_set (Set.Icc_subset_Icc_right aux4) + · exact fun k hk ↦ (intervalIntegrable_iff_integrableOn_Icc_of_le (mod_cast k.le_succ)).mpr + <| (integrablemulsum c ha hb hf_int).mono_set + <| (Set.Icc_subset_Icc_iff (mod_cast k.le_succ)).mpr <| mod_cast (ineqofmemIco hk) + +end abelSummationProof + +/-- Specialized version of `sum_mul_eq_sub_sub_integral_mul` for the case `a = 0`.-/ +theorem sum_mul_eq_sub_integral_mul {b : ℝ} (hb : 0 ≤ b) + (hf_diff : ∀ t ∈ Set.Icc 0 b, DifferentiableAt ℝ f t) + (hf_int : IntegrableOn (deriv f) (Set.Icc 0 b)) : + ∑ k ∈ Icc 0 ⌊b⌋₊, f k * c k = + f b * (∑ k ∈ Icc 0 ⌊b⌋₊, c k) - ∫ t in Set.Ioc 0 b, deriv f t * (∑ k ∈ Icc 0 ⌊t⌋₊, c k) := by + nth_rewrite 1 [Finset.Icc_eq_cons_Ioc (Nat.zero_le _)] + rw [sum_cons, ← Nat.floor_zero (α := ℝ), sum_mul_eq_sub_sub_integral_mul c le_rfl hb hf_diff + hf_int, Nat.floor_zero, Nat.cast_zero, Icc_self, sum_singleton] + ring + +/-- Specialized version of `sum_mul_eq_sub_integral_mul` when the first coefficient of the sequence +`c` is equal to `0`. -/ +theorem sum_mul_eq_sub_integral_mul' (hc : c 0 = 0) (b : ℝ) + (hf_diff : ∀ t ∈ Set.Icc 1 b, DifferentiableAt ℝ f t) + (hf_int : IntegrableOn (deriv f) (Set.Icc 1 b)) : + ∑ k ∈ Icc 0 ⌊b⌋₊, f k * c k = + f b * (∑ k ∈ Icc 0 ⌊b⌋₊, c k) - ∫ t in Set.Ioc 1 b, deriv f t * (∑ k ∈ Icc 0 ⌊t⌋₊, c k) := by + obtain hb | hb := le_or_gt 1 b + · have : 1 ≤ ⌊b⌋₊ := (Nat.one_le_floor_iff _).mpr hb + nth_rewrite 1 [Finset.Icc_eq_cons_Ioc (by linarith), sum_cons, ← Nat.Icc_succ_left, + Finset.Icc_eq_cons_Ioc (by linarith), sum_cons] + rw [Nat.succ_eq_add_one, zero_add, ← Nat.floor_one (α := ℝ), + sum_mul_eq_sub_sub_integral_mul c zero_le_one hb hf_diff hf_int, Nat.floor_one, Nat.cast_one, + Finset.Icc_eq_cons_Ioc zero_le_one, sum_cons, show 1 = 0 + 1 by rfl, Nat.Ioc_succ_singleton, + zero_add, sum_singleton, hc, mul_zero, zero_add] + ring + · simp_rw [Nat.floor_eq_zero.mpr hb, Icc_self, sum_singleton, Nat.cast_zero, hc, mul_zero, + Set.Ioc_eq_empty_of_le hb.le, Measure.restrict_empty, integral_zero_measure, sub_self] From e4876c14df1c7b54f95a1481526ebe1d580fc40f Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sun, 1 Dec 2024 09:59:44 +0000 Subject: [PATCH 444/829] chore: delete deprecated lemmas (#19627) This deletes a few deprecated lemmas, allowing us to clean up some imports. --- Mathlib/Algebra/Algebra/Equiv.lean | 22 ------------------ Mathlib/Algebra/Algebra/Hom.lean | 27 ----------------------- Mathlib/Algebra/Group/AddChar.lean | 2 ++ Mathlib/Algebra/TrivSqZeroExt.lean | 1 + Mathlib/Data/Matrix/ConjTranspose.lean | 1 - Mathlib/Data/Matrix/Notation.lean | 1 + Mathlib/RingTheory/Finiteness/Defs.lean | 1 + Mathlib/RingTheory/WittVector/IsPoly.lean | 3 +-- scripts/noshake.json | 3 ++- 9 files changed, 8 insertions(+), 53 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Equiv.lean b/Mathlib/Algebra/Algebra/Equiv.lean index cd98b733982f6..37da0ae421201 100644 --- a/Mathlib/Algebra/Algebra/Equiv.lean +++ b/Mathlib/Algebra/Algebra/Equiv.lean @@ -223,11 +223,6 @@ protected theorem map_smul (r : R) (x : A₁) : e (r • x) = r • e x := protected theorem map_pow : ∀ (x : A₁) (n : ℕ), e (x ^ n) = e x ^ n := map_pow _ -@[deprecated map_finsupp_sum (since := "2024-06-20")] -protected theorem map_finsupp_sum {α : Type*} [Zero α] {ι : Type*} (f : ι →₀ α) (g : ι → α → A₁) : - e (f.sum g) = f.sum fun i b => e (g i b) := - map_finsupp_sum _ _ _ - end map section bijective @@ -742,23 +737,6 @@ instance _root_.Finite.algEquiv [Finite (A₁ →ₐ[R] A₂)] : Finite (A₁ end Semiring -section CommSemiring - -variable [CommSemiring R] [CommSemiring A₁] [CommSemiring A₂] -variable [Algebra R A₁] [Algebra R A₂] (e : A₁ ≃ₐ[R] A₂) - -@[deprecated map_prod (since := "2024-06-20")] -protected theorem map_prod {ι : Type*} (f : ι → A₁) (s : Finset ι) : - e (∏ x ∈ s, f x) = ∏ x ∈ s, e (f x) := - map_prod _ f s - -@[deprecated map_finsupp_prod (since := "2024-06-20")] -protected theorem map_finsupp_prod {α : Type*} [Zero α] {ι : Type*} (f : ι →₀ α) (g : ι → α → A₁) : - e (f.prod g) = f.prod fun i a => e (g i a) := - map_finsupp_prod _ f g - -end CommSemiring - section Ring variable [CommSemiring R] [Ring A₁] [Ring A₂] diff --git a/Mathlib/Algebra/Algebra/Hom.lean b/Mathlib/Algebra/Algebra/Hom.lean index bdb71967c8292..2640de8265822 100644 --- a/Mathlib/Algebra/Algebra/Hom.lean +++ b/Mathlib/Algebra/Algebra/Hom.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Yury Kudryashov -/ import Mathlib.Algebra.Algebra.Basic -import Mathlib.Algebra.BigOperators.Finsupp /-! # Homomorphisms of `R`-algebras @@ -227,11 +226,6 @@ protected theorem map_sum {ι : Type*} (f : ι → A) (s : Finset ι) : φ (∑ x ∈ s, f x) = ∑ x ∈ s, φ (f x) := map_sum _ _ _ -@[deprecated map_finsupp_sum (since := "2024-06-26")] -protected theorem map_finsupp_sum {α : Type*} [Zero α] {ι : Type*} (f : ι →₀ α) (g : ι → α → A) : - φ (f.sum g) = f.sum fun i a => φ (g i a) := - map_finsupp_sum _ _ _ - /-- If a `RingHom` is `R`-linear, then it is an `AlgHom`. -/ def mk' (f : A →+* B) (h : ∀ (c : R) (x), f (c • x) = c • f x) : A →ₐ[R] B := { f with @@ -369,27 +363,6 @@ theorem algebraMap_eq_apply (f : A →ₐ[R] B) {y : R} {x : A} (h : algebraMap end Semiring -section CommSemiring - -variable [CommSemiring R] [CommSemiring A] [CommSemiring B] -variable [Algebra R A] [Algebra R B] (φ : A →ₐ[R] B) - -@[deprecated map_multiset_prod (since := "2024-06-26")] -protected theorem map_multiset_prod (s : Multiset A) : φ s.prod = (s.map φ).prod := - map_multiset_prod _ _ - -@[deprecated map_prod (since := "2024-06-26")] -protected theorem map_prod {ι : Type*} (f : ι → A) (s : Finset ι) : - φ (∏ x ∈ s, f x) = ∏ x ∈ s, φ (f x) := - map_prod _ _ _ - -@[deprecated map_finsupp_prod (since := "2024-06-26")] -protected theorem map_finsupp_prod {α : Type*} [Zero α] {ι : Type*} (f : ι →₀ α) (g : ι → α → A) : - φ (f.prod g) = f.prod fun i a => φ (g i a) := - map_finsupp_prod _ _ _ - -end CommSemiring - section Ring variable [CommSemiring R] [Ring A] [Ring B] diff --git a/Mathlib/Algebra/Group/AddChar.lean b/Mathlib/Algebra/Group/AddChar.lean index 0279bc188d39a..e8e36fe2f414f 100644 --- a/Mathlib/Algebra/Group/AddChar.lean +++ b/Mathlib/Algebra/Group/AddChar.lean @@ -5,6 +5,8 @@ Authors: Michael Stoll -/ import Mathlib.Algebra.Ring.Regular import Mathlib.Logic.Equiv.TransferInstance +import Mathlib.Algebra.BigOperators.Pi +import Mathlib.Algebra.BigOperators.Ring /-! # Characters from additive to multiplicative monoids diff --git a/Mathlib/Algebra/TrivSqZeroExt.lean b/Mathlib/Algebra/TrivSqZeroExt.lean index cc4239ab85e25..fd15145d858ba 100644 --- a/Mathlib/Algebra/TrivSqZeroExt.lean +++ b/Mathlib/Algebra/TrivSqZeroExt.lean @@ -6,6 +6,7 @@ Authors: Kenny Lau, Eric Wieser import Mathlib.Algebra.Algebra.Defs import Mathlib.Algebra.BigOperators.GroupWithZero.Action import Mathlib.LinearAlgebra.Prod +import Mathlib.Algebra.BigOperators.Pi /-! # Trivial Square-Zero Extension diff --git a/Mathlib/Data/Matrix/ConjTranspose.lean b/Mathlib/Data/Matrix/ConjTranspose.lean index 5aa7f18e50729..5f5843664ca6b 100644 --- a/Mathlib/Data/Matrix/ConjTranspose.lean +++ b/Mathlib/Data/Matrix/ConjTranspose.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Ellen Arlt, Blair Shi, Sean Leather, Mario Carneiro, Johan Commelin, Lu-Ming Zhang -/ import Mathlib.Algebra.BigOperators.GroupWithZero.Action -import Mathlib.Algebra.BigOperators.Pi import Mathlib.Algebra.BigOperators.Ring import Mathlib.Algebra.BigOperators.RingEquiv import Mathlib.Algebra.Module.Pi diff --git a/Mathlib/Data/Matrix/Notation.lean b/Mathlib/Data/Matrix/Notation.lean index 7ff272044152d..6d9f367504b17 100644 --- a/Mathlib/Data/Matrix/Notation.lean +++ b/Mathlib/Data/Matrix/Notation.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.Group.Fin.Tuple import Mathlib.Data.Matrix.RowCol import Mathlib.Data.Fin.VecNotation import Mathlib.Tactic.FinCases +import Mathlib.Algebra.BigOperators.Fin /-! # Matrix and vector notation diff --git a/Mathlib/RingTheory/Finiteness/Defs.lean b/Mathlib/RingTheory/Finiteness/Defs.lean index a56606fa8c8e9..fd8c45bb0be8d 100644 --- a/Mathlib/RingTheory/Finiteness/Defs.lean +++ b/Mathlib/RingTheory/Finiteness/Defs.lean @@ -5,6 +5,7 @@ Authors: Johan Commelin -/ import Mathlib.Algebra.Algebra.Hom import Mathlib.Data.Set.Finite.Lemmas +import Mathlib.Data.Finsupp.Defs import Mathlib.GroupTheory.Finiteness import Mathlib.RingTheory.Ideal.Span import Mathlib.Tactic.Algebraize diff --git a/Mathlib/RingTheory/WittVector/IsPoly.lean b/Mathlib/RingTheory/WittVector/IsPoly.lean index f6adc8664cb48..eed73bc325570 100644 --- a/Mathlib/RingTheory/WittVector/IsPoly.lean +++ b/Mathlib/RingTheory/WittVector/IsPoly.lean @@ -392,8 +392,7 @@ theorem map [Fact p.Prime] {f} (hf : IsPoly₂ p f) (g : R →+* S) (x y : 𝕎 end IsPoly₂ -attribute [ghost_simps] AlgHom.map_zero AlgHom.map_one AlgHom.map_add AlgHom.map_mul AlgHom.map_sub - AlgHom.map_neg AlgHom.id_apply map_natCast RingHom.map_zero RingHom.map_one RingHom.map_mul +attribute [ghost_simps] AlgHom.id_apply map_natCast RingHom.map_zero RingHom.map_one RingHom.map_mul RingHom.map_add RingHom.map_sub RingHom.map_neg RingHom.id_apply mul_add add_mul add_zero zero_add mul_one one_mul mul_zero zero_mul Nat.succ_ne_zero add_tsub_cancel_right Nat.succ_eq_add_one if_true eq_self_iff_true if_false forall_true_iff forall₂_true_iff diff --git a/scripts/noshake.json b/scripts/noshake.json index 2ca0558661bde..3126619078308 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -324,7 +324,8 @@ ["Mathlib.Algebra.MvPolynomial.CommRing", "Mathlib.Algebra.Polynomial.Basic"], "Mathlib.RingTheory.IntegralClosure.IsIntegral.Defs": ["Mathlib.Tactic.Algebraize"], - "Mathlib.RingTheory.Finiteness.Defs": ["Mathlib.Tactic.Algebraize"], + "Mathlib.RingTheory.Finiteness.Defs": + ["Mathlib.Data.Finsupp.Defs", "Mathlib.Tactic.Algebraize"], "Mathlib.RingTheory.Binomial": ["Mathlib.Algebra.Order.Floor"], "Mathlib.RingTheory.Adjoin.Basic": ["Mathlib.LinearAlgebra.Finsupp.SumProd"], "Mathlib.RepresentationTheory.FdRep": From 90f9e991a587838436180807998dbf129cf42db0 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sun, 1 Dec 2024 10:47:44 +0000 Subject: [PATCH 445/829] chore: remove PartENat from PowerSeries (#19622) --- Mathlib/Data/ENat/Basic.lean | 29 ++++ Mathlib/RingTheory/PowerSeries/Basic.lean | 4 + Mathlib/RingTheory/PowerSeries/Inverse.lean | 8 +- Mathlib/RingTheory/PowerSeries/Order.lean | 135 ++++++++---------- .../UniqueFactorizationDomain/Nat.lean | 2 +- 5 files changed, 97 insertions(+), 81 deletions(-) diff --git a/Mathlib/Data/ENat/Basic.lean b/Mathlib/Data/ENat/Basic.lean index f2e67d781c703..5b96e88c13f2d 100644 --- a/Mathlib/Data/ENat/Basic.lean +++ b/Mathlib/Data/ENat/Basic.lean @@ -59,6 +59,8 @@ variable {a b c m n : ℕ∞} `ℕ → ℕ∞` is `Nat.cast`. -/ @[simp] theorem some_eq_coe : (WithTop.some : ℕ → ℕ∞) = Nat.cast := rfl +theorem coe_inj {a b : ℕ} : (a : ℕ∞) = b ↔ a = b := WithTop.coe_inj + instance : SuccAddOrder ℕ∞ where succ_eq_add_one x := by cases x <;> simp [SuccOrder.succ] @@ -88,6 +90,29 @@ theorem coe_sub (m n : ℕ) : ↑(m - n) = (m - n : ℕ∞) := theorem top_pow {n : ℕ} (n_pos : 0 < n) : (⊤ : ℕ∞) ^ n = ⊤ := WithTop.top_pow n_pos +/-- Convert a `ℕ∞` to a `ℕ` using a proof that it is not infinite. -/ +def lift (x : ℕ∞) (h : x < ⊤) : ℕ := WithTop.untop x (WithTop.lt_top_iff_ne_top.mp h) + +@[simp] theorem coe_lift (x : ℕ∞) (h : x < ⊤) : (lift x h : ℕ∞) = x := + WithTop.coe_untop x (WithTop.lt_top_iff_ne_top.mp h) +@[simp] theorem lift_coe (n : ℕ) : lift (n : ℕ∞) (WithTop.coe_lt_top n) = n := rfl +@[simp] theorem lift_lt_iff {x : ℕ∞} {h} {n : ℕ} : lift x h < n ↔ x < n := WithTop.untop_lt_iff _ +@[simp] theorem lift_le_iff {x : ℕ∞} {h} {n : ℕ} : lift x h ≤ n ↔ x ≤ n := WithTop.untop_le_iff _ +@[simp] theorem lt_lift_iff {x : ℕ} {n : ℕ∞} {h} : x < lift n h ↔ x < n := WithTop.lt_untop_iff _ +@[simp] theorem le_lift_iff {x : ℕ} {n : ℕ∞} {h} : x ≤ lift n h ↔ x ≤ n := WithTop.le_untop_iff _ + +@[simp] theorem lift_zero : lift 0 (WithTop.coe_lt_top 0) = 0 := rfl +@[simp] theorem lift_one : lift 1 (WithTop.coe_lt_top 1) = 1 := rfl +@[simp] theorem lift_ofNat (n : ℕ) [n.AtLeastTwo] : + lift (no_index (OfNat.ofNat n)) (WithTop.coe_lt_top n) = OfNat.ofNat n := rfl + +@[simp] theorem add_lt_top {a b : ℕ∞} : a + b < ⊤ ↔ a < ⊤ ∧ b < ⊤ := WithTop.add_lt_top + +@[simp] theorem lift_add (a b : ℕ∞) (h : a + b < ⊤) : + lift (a + b) h = lift a (add_lt_top.1 h).1 + lift b (add_lt_top.1 h).2 := by + apply coe_inj.1 + simp + instance canLift : CanLift ℕ∞ ℕ (↑) (· ≠ ⊤) := WithTop.canLift instance : WellFoundedRelation ℕ∞ where @@ -262,6 +287,10 @@ lemma not_lt_zero (n : ℕ∞) : ¬ n < 0 := by lemma coe_lt_top (n : ℕ) : (n : ℕ∞) < ⊤ := WithTop.coe_lt_top n +lemma coe_lt_coe {n m : ℕ} : (n : ℕ∞) < (m : ℕ∞) ↔ n < m := by simp + +lemma coe_le_coe {n m : ℕ} : (n : ℕ∞) ≤ (m : ℕ∞) ↔ n ≤ m := by simp + @[elab_as_elim] theorem nat_induction {P : ℕ∞ → Prop} (a : ℕ∞) (h0 : P 0) (hsuc : ∀ n : ℕ, P n → P n.succ) (htop : (∀ n : ℕ, P n) → P ⊤) : P a := by diff --git a/Mathlib/RingTheory/PowerSeries/Basic.lean b/Mathlib/RingTheory/PowerSeries/Basic.lean index 85884f5a24154..2d18bb322885f 100644 --- a/Mathlib/RingTheory/PowerSeries/Basic.lean +++ b/Mathlib/RingTheory/PowerSeries/Basic.lean @@ -155,6 +155,10 @@ theorem ext {φ ψ : R⟦X⟧} (h : ∀ n, coeff R n φ = coeff R n ψ) : φ = · apply h rfl +@[simp] +theorem forall_coeff_eq_zero (φ : R⟦X⟧) : (∀ n, coeff R n φ = 0) ↔ φ = 0 := + ⟨fun h => ext h, fun h => by simp [h]⟩ + /-- Two formal power series are equal if all their coefficients are equal. -/ add_decl_doc PowerSeries.ext_iff diff --git a/Mathlib/RingTheory/PowerSeries/Inverse.lean b/Mathlib/RingTheory/PowerSeries/Inverse.lean index f69014616a2de..419143bde5cd0 100644 --- a/Mathlib/RingTheory/PowerSeries/Inverse.lean +++ b/Mathlib/RingTheory/PowerSeries/Inverse.lean @@ -202,7 +202,7 @@ theorem smul_inv (r : k) (φ : k⟦X⟧) : (r • φ)⁻¹ = r⁻¹ • φ⁻¹ /-- `firstUnitCoeff` is the non-zero coefficient whose index is `f.order`, seen as a unit of the field. It is obtained using `divided_by_X_pow_order`, defined in `PowerSeries.Order`-/ def firstUnitCoeff {f : k⟦X⟧} (hf : f ≠ 0) : kˣ := - let d := f.order.get (order_finite_iff_ne_zero.mpr hf) + let d := f.order.lift (order_finite_iff_ne_zero.mpr hf) have f_const : coeff k d f ≠ 0 := by apply coeff_order have : Invertible (constantCoeff k (divided_by_X_pow_order hf)) := by apply invertibleOfNonzero @@ -258,8 +258,8 @@ theorem Unit_of_divided_by_X_pow_order_zero : Unit_of_divided_by_X_pow_order (0 theorem eq_divided_by_X_pow_order_Iff_Unit {f : k⟦X⟧} (hf : f ≠ 0) : f = divided_by_X_pow_order hf ↔ IsUnit f := ⟨fun h ↦ by rw [h]; exact isUnit_divided_by_X_pow_order hf, fun h ↦ by - have : f.order.get (order_finite_iff_ne_zero.mpr hf) = 0 := by - simp only [order_zero_of_unit h, PartENat.get_zero] + have : f.order.lift (order_finite_iff_ne_zero.mpr hf) = 0 := by + simp [order_zero_of_unit h] convert (self_eq_X_pow_order_mul_divided_by_X_pow_order hf).symm simp only [this, pow_zero, one_mul]⟩ @@ -295,7 +295,7 @@ theorem hasUnitMulPowIrreducibleFactorization : ⟨X, And.intro X_irreducible (by intro f hf - use f.order.get (order_finite_iff_ne_zero.mpr hf) + use f.order.lift (order_finite_iff_ne_zero.mpr hf) use Unit_of_divided_by_X_pow_order f simp only [Unit_of_divided_by_X_pow_order_nonzero hf] exact self_eq_X_pow_order_mul_divided_by_X_pow_order hf)⟩ diff --git a/Mathlib/RingTheory/PowerSeries/Order.lean b/Mathlib/RingTheory/PowerSeries/Order.lean index d428546cd604a..70348c70fd666 100644 --- a/Mathlib/RingTheory/PowerSeries/Order.lean +++ b/Mathlib/RingTheory/PowerSeries/Order.lean @@ -7,7 +7,6 @@ Authors: Johan Commelin, Kenny Lau import Mathlib.Algebra.CharP.Defs import Mathlib.RingTheory.Multiplicity import Mathlib.RingTheory.PowerSeries.Basic -import Mathlib.Data.Nat.PartENat /-! # Formal power series (in one variable) - Order @@ -46,12 +45,11 @@ variable [Semiring R] {φ : R⟦X⟧} theorem exists_coeff_ne_zero_iff_ne_zero : (∃ n : ℕ, coeff R n φ ≠ 0) ↔ φ ≠ 0 := by refine not_iff_not.mp ?_ push_neg - -- FIXME: the `FunLike.coe` doesn't seem to be picked up in the expression after https://github.com/leanprover-community/mathlib4/pull/8386? - simp [PowerSeries.ext_iff, (coeff R _).map_zero] + simp [(coeff R _).map_zero] /-- The order of a formal power series `φ` is the greatest `n : PartENat` such that `X^n` divides `φ`. The order is `⊤` if and only if `φ = 0`. -/ -def order (φ : R⟦X⟧) : PartENat := +def order (φ : R⟦X⟧) : ℕ∞ := letI := Classical.decEq R letI := Classical.decEq R⟦X⟧ if h : φ = 0 then ⊤ else Nat.find (exists_coeff_ne_zero_iff_ne_zero.mpr h) @@ -61,20 +59,20 @@ def order (φ : R⟦X⟧) : PartENat := theorem order_zero : order (0 : R⟦X⟧) = ⊤ := dif_pos rfl -theorem order_finite_iff_ne_zero : (order φ).Dom ↔ φ ≠ 0 := by +theorem order_finite_iff_ne_zero : (order φ < ⊤) ↔ φ ≠ 0 := by simp only [order] constructor · split_ifs with h <;> intro H - · simp only [PartENat.top_eq_none, Part.not_none_dom] at H + · simp at H · exact h · intro h simp [h] /-- If the order of a formal power series is finite, then the coefficient indexed by the order is nonzero. -/ -theorem coeff_order (h : (order φ).Dom) : coeff R (φ.order.get h) φ ≠ 0 := by +theorem coeff_order (h : order φ < ⊤) : coeff R (φ.order.lift h) φ ≠ 0 := by classical - simp only [order, order_finite_iff_ne_zero.mp h, not_false_iff, dif_neg, PartENat.get_natCast'] + simp only [order, order_finite_iff_ne_zero.mp h, not_false_iff, dif_neg] generalize_proofs h exact Nat.find_spec h @@ -83,8 +81,7 @@ then the order of the power series is less than or equal to `n`. -/ theorem order_le (n : ℕ) (h : coeff R n φ ≠ 0) : order φ ≤ n := by classical rw [order, dif_neg] - · simp only [PartENat.coe_le_coe] - exact Nat.find_le h + · simpa using ⟨n, le_rfl, h⟩ · exact exists_coeff_ne_zero_iff_ne_zero.mp ⟨n, h⟩ /-- The `n`th coefficient of a formal power series is `0` if `n` is strictly @@ -95,28 +92,26 @@ theorem coeff_of_lt_order (n : ℕ) (h : ↑n < order φ) : coeff R n φ = 0 := /-- The `0` power series is the unique power series with infinite order. -/ @[simp] -theorem order_eq_top {φ : R⟦X⟧} : φ.order = ⊤ ↔ φ = 0 := - PartENat.not_dom_iff_eq_top.symm.trans order_finite_iff_ne_zero.not_left +theorem order_eq_top {φ : R⟦X⟧} : φ.order = ⊤ ↔ φ = 0 := by + simpa using order_finite_iff_ne_zero.not_left /-- The order of a formal power series is at least `n` if the `i`th coefficient is `0` for all `i < n`. -/ theorem nat_le_order (φ : R⟦X⟧) (n : ℕ) (h : ∀ i < n, coeff R i φ = 0) : ↑n ≤ order φ := by by_contra H; rw [not_le] at H - have : (order φ).Dom := PartENat.dom_of_le_natCast H.le - rw [← PartENat.natCast_get this, PartENat.coe_lt_coe] at H - exact coeff_order this (h _ H) + have lt_top : order φ < ⊤ := lt_top_of_lt H + replace H : (order φ).lift lt_top < n := by simpa + exact coeff_order lt_top (h _ H) /-- The order of a formal power series is at least `n` if the `i`th coefficient is `0` for all `i < n`. -/ -theorem le_order (φ : R⟦X⟧) (n : PartENat) (h : ∀ i : ℕ, ↑i < n → coeff R i φ = 0) : +theorem le_order (φ : R⟦X⟧) (n : ℕ∞) (h : ∀ i : ℕ, ↑i < n → coeff R i φ = 0) : n ≤ order φ := by - induction n using PartENat.casesOn - · show _ ≤ _ - rw [top_le_iff, order_eq_top] - ext i - exact h _ (PartENat.natCast_lt_top i) - · apply nat_le_order - simpa only [PartENat.coe_lt_coe] using h + cases n with + | top => simpa using ext (by simpa using h) + | coe n => + convert nat_le_order φ n _ + simpa using h /-- The order of a formal power series is exactly `n` if the `n`th coefficient is nonzero, and the `i`th coefficient is `0` for all `i < n`. -/ @@ -124,25 +119,17 @@ theorem order_eq_nat {φ : R⟦X⟧} {n : ℕ} : order φ = n ↔ coeff R n φ ≠ 0 ∧ ∀ i, i < n → coeff R i φ = 0 := by classical rcases eq_or_ne φ 0 with (rfl | hφ) - · simpa [(coeff R _).map_zero] using (PartENat.natCast_ne_top _).symm + · simp simp [order, dif_neg hφ, Nat.find_eq_iff] /-- The order of a formal power series is exactly `n` if the `n`th coefficient is nonzero, and the `i`th coefficient is `0` for all `i < n`. -/ -theorem order_eq {φ : R⟦X⟧} {n : PartENat} : +theorem order_eq {φ : R⟦X⟧} {n : ℕ∞} : order φ = n ↔ (∀ i : ℕ, ↑i = n → coeff R i φ ≠ 0) ∧ ∀ i : ℕ, ↑i < n → coeff R i φ = 0 := by - induction n using PartENat.casesOn - · rw [order_eq_top] - constructor - · rintro rfl - constructor <;> intros - · exfalso - exact PartENat.natCast_ne_top ‹_› ‹_› - · exact (coeff _ _).map_zero - · rintro ⟨_h₁, h₂⟩ - ext i - exact h₂ i (PartENat.natCast_lt_top i) - · simpa [PartENat.natCast_inj] using order_eq_nat + cases n with + | top => simp [ext_iff] + | coe n => simp [order_eq_nat] + /-- The order of the sum of two formal power series is at least the minimum of their orders. -/ @@ -197,14 +184,14 @@ alias order_mul_ge := le_order_mul /-- The order of the monomial `a*X^n` is infinite if `a = 0` and `n` otherwise. -/ theorem order_monomial (n : ℕ) (a : R) [Decidable (a = 0)] : - order (monomial R n a) = if a = 0 then (⊤ : PartENat) else n := by + order (monomial R n a) = if a = 0 then (⊤ : ℕ∞) else n := by split_ifs with h · rw [h, order_eq_top, LinearMap.map_zero] · rw [order_eq] constructor <;> intro i hi - · rw [PartENat.natCast_inj] at hi + · simp only [Nat.cast_inj] at hi rwa [hi, coeff_monomial_same] - · rw [PartENat.coe_lt_coe] at hi + · simp only [Nat.cast_lt] at hi rw [coeff_monomial, if_neg] exact ne_of_lt hi @@ -242,42 +229,42 @@ theorem coeff_mul_prod_one_sub_of_lt_order {R ι : Type*} [CommRing R] (k : ℕ) exact ih t.2 -- TODO: link with `X_pow_dvd_iff` -theorem X_pow_order_dvd (h : (order φ).Dom) : X ^ (order φ).get h ∣ φ := by - refine ⟨PowerSeries.mk fun n => coeff R (n + (order φ).get h) φ, ?_⟩ +theorem X_pow_order_dvd (h : order φ < ⊤) : X ^ (order φ).lift h ∣ φ := by + refine ⟨PowerSeries.mk fun n => coeff R (n + (order φ).lift h) φ, ?_⟩ ext n simp only [coeff_mul, coeff_X_pow, coeff_mk, boole_mul, Finset.sum_ite, Finset.sum_const_zero, add_zero] - rw [Finset.filter_fst_eq_antidiagonal n (Part.get (order φ) h)] + rw [Finset.filter_fst_eq_antidiagonal n ((order φ).lift h)] split_ifs with hn · simp [tsub_add_cancel_of_le hn] · simp only [Finset.sum_empty] refine coeff_of_lt_order _ ?_ - simpa [PartENat.coe_lt_iff] using fun _ => hn + simpa using hn theorem order_eq_emultiplicity_X {R : Type*} [Semiring R] (φ : R⟦X⟧) : order φ = emultiplicity X φ := by classical rcases eq_or_ne φ 0 with (rfl | hφ) · simp - induction' ho : order φ using PartENat.casesOn with n - · simp [hφ] at ho - have hn : φ.order.get (order_finite_iff_ne_zero.mpr hφ) = n := by simp [ho] - rw [← hn, ← PartENat.ofENat_coe, eq_comm] - congr 1 - apply le_antisymm _ - · apply le_emultiplicity_of_pow_dvd - apply X_pow_order_dvd - · apply Order.le_of_lt_add_one - rw [← not_le, ← Nat.cast_one, ← Nat.cast_add, ← pow_dvd_iff_le_emultiplicity] - rintro ⟨ψ, H⟩ - have := congr_arg (coeff R n) H - rw [← (ψ.commute_X.pow_right _).eq, coeff_mul_of_lt_order, ← hn] at this - · exact coeff_order _ this - · rw [X_pow_eq, order_monomial] - split_ifs - · exact PartENat.natCast_lt_top _ - · rw [← hn, PartENat.coe_lt_coe] - exact Nat.lt_succ_self _ + cases ho : order φ with + | top => simp [hφ] at ho + | coe n => + have hn : φ.order.lift (order_finite_iff_ne_zero.mpr hφ) = n := by simp [ho] + rw [← hn, eq_comm] + apply le_antisymm _ + · apply le_emultiplicity_of_pow_dvd + apply X_pow_order_dvd + · apply Order.le_of_lt_add_one + rw [← not_le, ← Nat.cast_one, ← Nat.cast_add, ← pow_dvd_iff_le_emultiplicity] + rintro ⟨ψ, H⟩ + have := congr_arg (coeff R n) H + rw [← (ψ.commute_X.pow_right _).eq, coeff_mul_of_lt_order, ← hn] at this + · exact coeff_order _ this + · rw [X_pow_eq, order_monomial] + split_ifs + · simp + · rw [← hn, ENat.coe_lt_coe] + simp /-- Given a non-zero power series `f`, `divided_by_X_pow_order f` is the power series obtained by dividing out the largest power of X that divides `f`, that is its order -/ @@ -285,7 +272,7 @@ def divided_by_X_pow_order {f : PowerSeries R} (hf : f ≠ 0) : R⟦X⟧ := (exists_eq_mul_right_of_dvd (X_pow_order_dvd (order_finite_iff_ne_zero.2 hf))).choose theorem self_eq_X_pow_order_mul_divided_by_X_pow_order {f : R⟦X⟧} (hf : f ≠ 0) : - X ^ f.order.get (order_finite_iff_ne_zero.mpr hf) * divided_by_X_pow_order hf = f := + X ^ f.order.lift (order_finite_iff_ne_zero.mpr hf) * divided_by_X_pow_order hf = f := haveI dvd := X_pow_order_dvd (order_finite_iff_ne_zero.mpr hf) (exists_eq_mul_right_of_dvd dvd).choose_spec.symm @@ -329,28 +316,24 @@ variable [CommRing R] [IsDomain R] is the sum of their orders. -/ theorem order_mul (φ ψ : R⟦X⟧) : order (φ * ψ) = order φ + order ψ := by classical - simp_rw [order_eq_emultiplicity_X] - change PartENat.withTopAddEquiv.symm _ = - PartENat.withTopAddEquiv.symm _ + PartENat.withTopAddEquiv.symm _ - rw [← map_add] - congr 1 - exact emultiplicity_mul X_prime + simp only [order_eq_emultiplicity_X] + rw [emultiplicity_mul X_prime] -- Dividing `X` by the maximal power of `X` dividing it leaves `1`. @[simp] theorem divided_by_X_pow_order_of_X_eq_one : divided_by_X_pow_order X_ne_zero = (1 : R⟦X⟧) := by rw [← mul_eq_left₀ X_ne_zero] - simpa only [order_X, X_ne_zero, PartENat.get_one, pow_one, Ne, not_false_iff] using - self_eq_X_pow_order_mul_divided_by_X_pow_order (@X_ne_zero R _ _) + simpa using self_eq_X_pow_order_mul_divided_by_X_pow_order (@X_ne_zero R _ _) -- Dividing a power series by the maximal power of `X` dividing it, respects multiplication. theorem divided_by_X_pow_orderMul {f g : R⟦X⟧} (hf : f ≠ 0) (hg : g ≠ 0) : divided_by_X_pow_order hf * divided_by_X_pow_order hg = divided_by_X_pow_order (mul_ne_zero hf hg) := by - set df := f.order.get (order_finite_iff_ne_zero.mpr hf) - set dg := g.order.get (order_finite_iff_ne_zero.mpr hg) - set dfg := (f * g).order.get (order_finite_iff_ne_zero.mpr (mul_ne_zero hf hg)) with hdfg - have H_add_d : df + dg = dfg := by simp_all only [PartENat.get_add, order_mul f g] + set df := f.order.lift (order_finite_iff_ne_zero.mpr hf) + set dg := g.order.lift (order_finite_iff_ne_zero.mpr hg) + set dfg := (f * g).order.lift (order_finite_iff_ne_zero.mpr (mul_ne_zero hf hg)) with hdfg + have H_add_d : df + dg = dfg := by + simp_all [order_mul f g] have H := self_eq_X_pow_order_mul_divided_by_X_pow_order (mul_ne_zero hf hg) have : f * g = X ^ dfg * (divided_by_X_pow_order hf * divided_by_X_pow_order hg) := by calc diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean index 4c69bfb655b75..d8e8cd87f6ef9 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean @@ -27,7 +27,7 @@ instance instWfDvdMonoid : WfDvdMonoid ℕ where revert h simp [DvdNotUnit] cases b - · simpa [succ_ne_zero] using WithTop.coe_lt_top (a + 1) + · simpa [succ_ne_zero] using ENat.coe_lt_top (a + 1) cases' dvd_and_not_dvd_iff.2 h with h1 h2 simp only [succ_ne_zero, cast_lt, if_false] refine lt_of_le_of_ne (Nat.le_of_dvd (Nat.succ_pos _) h1) fun con => h2 ?_ From a3fcd535e9295e67a5a9d10eef9ed013a5d90f12 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Sun, 1 Dec 2024 11:18:33 +0000 Subject: [PATCH 446/829] =?UTF-8?q?chore(Algebra/Data/Order/Floor):=20remo?= =?UTF-8?q?ve=20todo=20on=20`fract=5Fdiv=5FintCast=5Feq=E2=80=A6=20(#19652?= =?UTF-8?q?)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit …_div_intCast_mod` There are no plans to implement this, because of there are no plans to add more API for `Int.fmod`. Co-authored-by: Moritz Firsching --- Mathlib/Algebra/Order/Floor.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Mathlib/Algebra/Order/Floor.lean b/Mathlib/Algebra/Order/Floor.lean index 0ec2dc284de96..286ee58459e4b 100644 --- a/Mathlib/Algebra/Order/Floor.lean +++ b/Mathlib/Algebra/Order/Floor.lean @@ -1044,7 +1044,6 @@ theorem fract_div_natCast_eq_div_natCast_mod {m n : ℕ} : fract ((m : k) / n) = norm_cast rw [← Nat.cast_add, Nat.mod_add_div m n] --- TODO Generalise this to allow `n : ℤ` using `Int.fmod` instead of `Int.mod`. theorem fract_div_intCast_eq_div_intCast_mod {m : ℤ} {n : ℕ} : fract ((m : k) / n) = ↑(m % n) / n := by rcases n.eq_zero_or_pos with (rfl | hn) From 9d16cdd2432b261c1ae3bff386c623a58ccf7c40 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sun, 1 Dec 2024 13:50:00 +0000 Subject: [PATCH 447/829] chore: update benchmark runner for leanprover/lean4#5684 (#19656) Trying again, this time using [#18354](https://github.com/leanprover-community/mathlib4/pull/18354). --- scripts/bench/temci-config.run.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/scripts/bench/temci-config.run.yml b/scripts/bench/temci-config.run.yml index 49179c7e7dad2..0d3692b08a9fb 100644 --- a/scripts/bench/temci-config.run.yml +++ b/scripts/bench/temci-config.run.yml @@ -7,7 +7,7 @@ rusage_properties: ['maxrss'] cmd: | # use build cache for proofwidgets, but not for anything else - bash -c 'set -eo pipefail; lake clean 1>&2 && LEAN_PATH=$(lean --print-libdir) lake build proofwidgets 1>&2 && rm -f .lake/packages/batteries/.lake/build/bin/runLinter 1>&2 && LEAN=./scripts/bench/fake-root/bin/lean lake build --no-cache -v | ./scripts/bench/accumulate_profile.py | grep -v took' + bash -c 'set -eo pipefail; lake clean 1>&2 && LEAN_PATH=$(lean --print-libdir) lake build proofwidgets 1>&2 && rm -f .lake/packages/batteries/.lake/build/bin/runLinter 1>&2 && LAKE_OVERRIDE_LEAN=true LEAN=$(readlink -m ./scripts/bench/fake-root/bin/lean) lake build --no-cache -v | ./scripts/bench/accumulate_profile.py | grep -v took' parse_output: true runs: 1 - attributes: From 65cc54340c823eff8861b7d448c2755bbc87bae8 Mon Sep 17 00:00:00 2001 From: Bryan Gin-ge Chen Date: Sun, 1 Dec 2024 14:12:31 +0000 Subject: [PATCH 448/829] chore: add location of build.in.yml to directories checked by dependabot (#18712) Not sure if this will work, but since dependabot currently doesn't try to update `build.in.yml`, our linter rejects its PRs. cf. #18026 --- .github/dependabot.yml | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/.github/dependabot.yml b/.github/dependabot.yml index 4e43547be01ea..c3c49172d1f27 100644 --- a/.github/dependabot.yml +++ b/.github/dependabot.yml @@ -3,7 +3,9 @@ version: 2 # Specifies the version of the Dependabot configuration file format updates: # Configuration for dependency updates - package-ecosystem: "github-actions" # Specifies the ecosystem to check for updates - directory: "/" # Specifies the directory to check for dependencies; "/" means the root directory + directories: + - "/" # Specifies the directory to check for dependencies; "/" means the root directory + - "/.github/" # covers `build.in.yml` as well, which is not in `.github/workflows/` because it shouldn't be run in CI. schedule: # Check for updates to GitHub Actions every month interval: "monthly" From 2aefce62d6d933585e9c688b080ebeadc5c8e7ee Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 1 Dec 2024 16:15:39 +0000 Subject: [PATCH 449/829] chore: move `Topology.Support` to `Topology.Algebra.Support` (#19651) to make it clear it's not pure topology. Indeed, the support is the preimage of `0`, which is an algebraic object. Sounds quite trivial like that, but the file actually imports `Module`! --- Mathlib.lean | 2 +- Mathlib/Topology/Algebra/ConstMulAction.lean | 2 +- Mathlib/Topology/{ => Algebra}/Support.lean | 0 Mathlib/Topology/ContinuousMap/CompactlySupported.lean | 2 +- Mathlib/Topology/Homeomorph.lean | 4 ++-- Mathlib/Topology/MetricSpace/ProperSpace.lean | 4 ++-- Mathlib/Topology/Order/Compact.lean | 4 ++-- Mathlib/Topology/UniformSpace/HeineCantor.lean | 4 ++-- 8 files changed, 11 insertions(+), 11 deletions(-) rename Mathlib/Topology/{ => Algebra}/Support.lean (100%) diff --git a/Mathlib.lean b/Mathlib.lean index f120c15dfeb9e..99e0381f97d9a 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4994,6 +4994,7 @@ import Mathlib.Topology.Algebra.SeparationQuotient.Basic import Mathlib.Topology.Algebra.SeparationQuotient.Section import Mathlib.Topology.Algebra.Star import Mathlib.Topology.Algebra.StarSubalgebra +import Mathlib.Topology.Algebra.Support import Mathlib.Topology.Algebra.UniformConvergence import Mathlib.Topology.Algebra.UniformField import Mathlib.Topology.Algebra.UniformFilterBasis @@ -5326,7 +5327,6 @@ import Mathlib.Topology.Sober import Mathlib.Topology.Specialization import Mathlib.Topology.Spectral.Hom import Mathlib.Topology.StoneCech -import Mathlib.Topology.Support import Mathlib.Topology.TietzeExtension import Mathlib.Topology.Ultrafilter import Mathlib.Topology.UniformSpace.AbsoluteValue diff --git a/Mathlib/Topology/Algebra/ConstMulAction.lean b/Mathlib/Topology/Algebra/ConstMulAction.lean index 10007257a47c5..cd4bdb42717b6 100644 --- a/Mathlib/Topology/Algebra/ConstMulAction.lean +++ b/Mathlib/Topology/Algebra/ConstMulAction.lean @@ -8,9 +8,9 @@ import Mathlib.Algebra.Order.Group.Synonym import Mathlib.Data.Set.Pointwise.SMul import Mathlib.GroupTheory.GroupAction.Defs import Mathlib.Topology.Algebra.Constructions +import Mathlib.Topology.Algebra.Support import Mathlib.Topology.Bases import Mathlib.Topology.Homeomorph -import Mathlib.Topology.Support /-! # Monoid actions continuous in the second variable diff --git a/Mathlib/Topology/Support.lean b/Mathlib/Topology/Algebra/Support.lean similarity index 100% rename from Mathlib/Topology/Support.lean rename to Mathlib/Topology/Algebra/Support.lean diff --git a/Mathlib/Topology/ContinuousMap/CompactlySupported.lean b/Mathlib/Topology/ContinuousMap/CompactlySupported.lean index cb59956edbce0..eb5348c9f448c 100644 --- a/Mathlib/Topology/ContinuousMap/CompactlySupported.lean +++ b/Mathlib/Topology/ContinuousMap/CompactlySupported.lean @@ -3,9 +3,9 @@ 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.Algebra.Support import Mathlib.Topology.ContinuousMap.CocompactMap import Mathlib.Topology.ContinuousMap.ZeroAtInfty -import Mathlib.Topology.Support /-! # Compactly supported continuous functions diff --git a/Mathlib/Topology/Homeomorph.lean b/Mathlib/Topology/Homeomorph.lean index f536c0c7e30dc..f98a10c4d8b6b 100644 --- a/Mathlib/Topology/Homeomorph.lean +++ b/Mathlib/Topology/Homeomorph.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Patrick Massot, Sébastien Gouëzel, Zhouhang Zhou, Reid Barton -/ import Mathlib.Logic.Equiv.Fin -import Mathlib.Topology.DenseEmbedding -import Mathlib.Topology.Support +import Mathlib.Topology.Algebra.Support import Mathlib.Topology.Connected.LocallyConnected +import Mathlib.Topology.DenseEmbedding /-! # Homeomorphisms diff --git a/Mathlib/Topology/MetricSpace/ProperSpace.lean b/Mathlib/Topology/MetricSpace/ProperSpace.lean index 3868c4d9c6a97..69d91fcc8b806 100644 --- a/Mathlib/Topology/MetricSpace/ProperSpace.lean +++ b/Mathlib/Topology/MetricSpace/ProperSpace.lean @@ -4,11 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import Mathlib.Topology.Order.IsLUB -import Mathlib.Topology.Support +import Mathlib.Topology.Algebra.Support import Mathlib.Topology.MetricSpace.Pseudo.Basic import Mathlib.Topology.MetricSpace.Pseudo.Lemmas import Mathlib.Topology.MetricSpace.Pseudo.Pi +import Mathlib.Topology.Order.IsLUB /-! ## Proper spaces diff --git a/Mathlib/Topology/Order/Compact.lean b/Mathlib/Topology/Order/Compact.lean index bf30e3961311d..c8c4297800a48 100644 --- a/Mathlib/Topology/Order/Compact.lean +++ b/Mathlib/Topology/Order/Compact.lean @@ -3,10 +3,10 @@ Copyright (c) 2021 Patrick Massot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Yury Kudryashov -/ -import Mathlib.Topology.Order.LocalExtr +import Mathlib.Topology.Algebra.Support import Mathlib.Topology.Order.IntermediateValue -import Mathlib.Topology.Support import Mathlib.Topology.Order.IsLUB +import Mathlib.Topology.Order.LocalExtr /-! # Compactness of a closed interval diff --git a/Mathlib/Topology/UniformSpace/HeineCantor.lean b/Mathlib/Topology/UniformSpace/HeineCantor.lean index d97593eba0566..166e33321e049 100644 --- a/Mathlib/Topology/UniformSpace/HeineCantor.lean +++ b/Mathlib/Topology/UniformSpace/HeineCantor.lean @@ -3,9 +3,9 @@ Copyright (c) 2020 Patrick Massot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Yury Kudryashov -/ -import Mathlib.Topology.UniformSpace.Equicontinuity +import Mathlib.Topology.Algebra.Support import Mathlib.Topology.UniformSpace.Compact -import Mathlib.Topology.Support +import Mathlib.Topology.UniformSpace.Equicontinuity /-! # Compact separated uniform spaces From 6926cee95dfac950113a35dda8c99defd6fb5442 Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Sun, 1 Dec 2024 19:53:14 +0000 Subject: [PATCH 450/829] fix(FieldTheory/AlgebraicClosure): `algebraicClosure.AlgEquiv.algebraicClosure` -> `AlgEquiv.algebraicClosure` (#19642) --- Mathlib/FieldTheory/AlgebraicClosure.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/FieldTheory/AlgebraicClosure.lean b/Mathlib/FieldTheory/AlgebraicClosure.lean index 8307314cbe26b..9a4a0902bc605 100644 --- a/Mathlib/FieldTheory/AlgebraicClosure.lean +++ b/Mathlib/FieldTheory/AlgebraicClosure.lean @@ -94,7 +94,7 @@ def algEquivOfAlgEquiv (i : E ≃ₐ[F] K) : algebraicClosure F E ≃ₐ[F] algebraicClosure F K := (intermediateFieldMap i _).trans (equivOfEq (map_eq_of_algEquiv i)) -alias AlgEquiv.algebraicClosure := algebraicClosure.algEquivOfAlgEquiv +alias _root_.AlgEquiv.algebraicClosure := algEquivOfAlgEquiv variable (F E K) From 439bd5822664105871736635701caf0a02a962c7 Mon Sep 17 00:00:00 2001 From: Bryan Gin-ge Chen Date: Sun, 1 Dec 2024 21:58:55 +0000 Subject: [PATCH 451/829] ci(dependabot.yml): use glob syntax in config... (#19662) to try to get dependabot to create only one PR instead of several, cf. #19655, #19658 --- .github/dependabot.yml | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/.github/dependabot.yml b/.github/dependabot.yml index c3c49172d1f27..79181e47c2b08 100644 --- a/.github/dependabot.yml +++ b/.github/dependabot.yml @@ -4,8 +4,7 @@ updates: # Configuration for dependency updates - package-ecosystem: "github-actions" # Specifies the ecosystem to check for updates directories: - - "/" # Specifies the directory to check for dependencies; "/" means the root directory - - "/.github/" # covers `build.in.yml` as well, which is not in `.github/workflows/` because it shouldn't be run in CI. + - "/.github/*" # covers `build.in.yml` as well, which is not in `.github/workflows/` because it shouldn't be run in CI. schedule: # Check for updates to GitHub Actions every month interval: "monthly" From 8639aea31f527d1707614a8faf58611e649cdf93 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 1 Dec 2024 22:43:01 +0000 Subject: [PATCH 452/829] feat: `Polynomial.map f` strictly decreases the degree if `f p.leadingCoeff = 0` (#19411) Also make more arguments to `degree_map_le` implicit. From GrowthInGroups Co-authored-by: Andrew Yang --- Mathlib/Algebra/Polynomial/Degree/Lemmas.lean | 2 +- Mathlib/Algebra/Polynomial/Div.lean | 2 +- Mathlib/Algebra/Polynomial/Eval/Degree.lean | 54 ++++++++++++------- Mathlib/Algebra/Polynomial/Monic.lean | 4 +- Mathlib/Algebra/Polynomial/Roots.lean | 4 +- Mathlib/Algebra/Polynomial/Splits.lean | 4 +- .../Polynomial/SumIteratedDerivative.lean | 19 ++++--- .../PrimeSpectrum/Basic.lean | 2 +- Mathlib/FieldTheory/Minpoly/MinpolyDiv.lean | 2 +- Mathlib/RingTheory/EisensteinCriterion.lean | 4 +- .../RingTheory/MvPolynomial/Homogeneous.lean | 2 +- .../Polynomial/Eisenstein/Basic.lean | 4 +- Mathlib/RingTheory/RootsOfUnity/Minpoly.lean | 2 +- 13 files changed, 61 insertions(+), 44 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean b/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean index 22a6b53a366a2..5e474102b75cc 100644 --- a/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean +++ b/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean @@ -271,7 +271,7 @@ theorem degree_map_eq_iff {f : R →+* S} {p : Polynomial R} : theorem natDegree_map_eq_iff {f : R →+* S} {p : Polynomial R} : natDegree (map f p) = natDegree p ↔ f (p.leadingCoeff) ≠ 0 ∨ natDegree p = 0 := by rcases eq_or_ne (natDegree p) 0 with h|h - · simp_rw [h, ne_eq, or_true, iff_true, ← Nat.le_zero, ← h, natDegree_map_le f p] + · simp_rw [h, ne_eq, or_true, iff_true, ← Nat.le_zero, ← h, natDegree_map_le] have h2 : p ≠ 0 := by rintro rfl; simp at h have h3 : degree p ≠ (0 : ℕ) := degree_ne_of_natDegree_ne h simp_rw [h, or_false, natDegree, WithBot.unbot'_eq_unbot'_iff, degree_map_eq_iff] diff --git a/Mathlib/Algebra/Polynomial/Div.lean b/Mathlib/Algebra/Polynomial/Div.lean index 804749c1a188a..03aa336587516 100644 --- a/Mathlib/Algebra/Polynomial/Div.lean +++ b/Mathlib/Algebra/Polynomial/Div.lean @@ -349,7 +349,7 @@ theorem map_mod_divByMonic [Ring S] (f : R →+* S) (hq : Monic q) : div_modByMonic_unique ((p /ₘ q).map f) _ (hq.map f) ⟨Eq.symm <| by rw [← Polynomial.map_mul, ← Polynomial.map_add, modByMonic_add_div _ hq], calc - _ ≤ degree (p %ₘ q) := degree_map_le _ _ + _ ≤ degree (p %ₘ q) := degree_map_le _ < degree q := degree_modByMonic_lt _ hq _ = _ := Eq.symm <| diff --git a/Mathlib/Algebra/Polynomial/Eval/Degree.lean b/Mathlib/Algebra/Polynomial/Eval/Degree.lean index 713aca85943c1..41edbc137ec86 100644 --- a/Mathlib/Algebra/Polynomial/Eval/Degree.lean +++ b/Mathlib/Algebra/Polynomial/Eval/Degree.lean @@ -111,25 +111,15 @@ end Comp section Map -variable [Semiring S] -variable (f : R →+* S) +variable [Semiring S] {f : R →+* S} {p : R[X]} +variable (f) in /-- If `R` and `S` are isomorphic, then so are their polynomial rings. -/ @[simps!] def mapEquiv (e : R ≃+* S) : R[X] ≃+* S[X] := RingEquiv.ofHomInv (mapRingHom (e : R →+* S)) (mapRingHom (e.symm : S →+* R)) (by ext; simp) (by ext; simp) -theorem degree_map_le (p : R[X]) : degree (p.map f) ≤ degree p := by - refine (degree_le_iff_coeff_zero _ _).2 fun m hm => ?_ - rw [degree_lt_iff_coeff_zero] at hm - simp [hm m le_rfl] - -theorem natDegree_map_le (p : R[X]) : natDegree (p.map f) ≤ natDegree p := - natDegree_le_natDegree (degree_map_le f p) - -variable {f} - theorem map_monic_eq_zero_iff (hp : p.Monic) : p.map f = 0 ↔ ∀ x, f x = 0 := ⟨fun hfp x => calc @@ -142,15 +132,39 @@ theorem map_monic_eq_zero_iff (hp : p.Monic) : p.map f = 0 ↔ ∀ x, f x = 0 := theorem map_monic_ne_zero (hp : p.Monic) [Nontrivial S] : p.map f ≠ 0 := fun h => f.map_one_ne_zero ((map_monic_eq_zero_iff hp).mp h _) +lemma degree_map_le : degree (p.map f) ≤ degree p := by + refine (degree_le_iff_coeff_zero _ _).2 fun m hm => ?_ + rw [degree_lt_iff_coeff_zero] at hm + simp [hm m le_rfl] + +lemma natDegree_map_le : natDegree (p.map f) ≤ natDegree p := natDegree_le_natDegree degree_map_le + +lemma degree_map_lt (hp : f p.leadingCoeff = 0) (hp₀ : p ≠ 0) : (p.map f).degree < p.degree := by + refine degree_map_le.lt_of_ne fun hpq ↦ hp₀ ?_ + rw [leadingCoeff, ← coeff_map, ← natDegree_eq_natDegree hpq, ← leadingCoeff, leadingCoeff_eq_zero] + at hp + rw [← degree_eq_bot, ← hpq, hp, degree_zero] + +lemma natDegree_map_lt (hp : f p.leadingCoeff = 0) (hp₀ : map f p ≠ 0) : + (p.map f).natDegree < p.natDegree := + natDegree_lt_natDegree hp₀ <| degree_map_lt hp <| by rintro rfl; simp at hp₀ + +/-- Variant of `natDegree_map_lt` that assumes `0 < natDegree p` instead of `map f p ≠ 0`. -/ +lemma natDegree_map_lt' (hp : f p.leadingCoeff = 0) (hp₀ : 0 < natDegree p) : + (p.map f).natDegree < p.natDegree := by + by_cases H : map f p = 0 + · rwa [H, natDegree_zero] + · exact natDegree_map_lt hp H + theorem degree_map_eq_of_leadingCoeff_ne_zero (f : R →+* S) (hf : f (leadingCoeff p) ≠ 0) : - degree (p.map f) = degree p := - le_antisymm (degree_map_le f _) <| by - have hp0 : p ≠ 0 := - leadingCoeff_ne_zero.mp fun hp0 => hf (_root_.trans (congr_arg _ hp0) f.map_zero) - rw [degree_eq_natDegree hp0] - refine le_degree_of_ne_zero ?_ - rw [coeff_map] - exact hf + degree (p.map f) = degree p := by + refine degree_map_le.antisymm ?_ + have hp0 : p ≠ 0 := + leadingCoeff_ne_zero.mp fun hp0 => hf (_root_.trans (congr_arg _ hp0) f.map_zero) + rw [degree_eq_natDegree hp0] + refine le_degree_of_ne_zero ?_ + rw [coeff_map] + exact hf theorem natDegree_map_of_leadingCoeff_ne_zero (f : R →+* S) (hf : f (leadingCoeff p) ≠ 0) : natDegree (p.map f) = natDegree p := diff --git a/Mathlib/Algebra/Polynomial/Monic.lean b/Mathlib/Algebra/Polynomial/Monic.lean index ee43be7d90efb..f78aa81f906ba 100644 --- a/Mathlib/Algebra/Polynomial/Monic.lean +++ b/Mathlib/Algebra/Polynomial/Monic.lean @@ -357,7 +357,7 @@ variable [Semiring R] @[simp] theorem Monic.natDegree_map [Semiring S] [Nontrivial S] {P : R[X]} (hmo : P.Monic) (f : R →+* S) : (P.map f).natDegree = P.natDegree := by - refine le_antisymm (natDegree_map_le _ _) (le_natDegree_of_ne_zero ?_) + refine le_antisymm natDegree_map_le (le_natDegree_of_ne_zero ?_) rw [coeff_map, Monic.coeff_natDegree hmo, RingHom.map_one] exact one_ne_zero @@ -366,7 +366,7 @@ theorem Monic.degree_map [Semiring S] [Nontrivial S] {P : R[X]} (hmo : P.Monic) (P.map f).degree = P.degree := by by_cases hP : P = 0 · simp [hP] - · refine le_antisymm (degree_map_le _ _) ?_ + · refine le_antisymm degree_map_le ?_ rw [degree_eq_natDegree hP] refine le_degree_of_ne_zero ?_ rw [coeff_map, Monic.coeff_natDegree hmo, RingHom.map_one] diff --git a/Mathlib/Algebra/Polynomial/Roots.lean b/Mathlib/Algebra/Polynomial/Roots.lean index 2078ef2aeeb5f..758a417dfb770 100644 --- a/Mathlib/Algebra/Polynomial/Roots.lean +++ b/Mathlib/Algebra/Polynomial/Roots.lean @@ -749,13 +749,13 @@ theorem roots_map_of_injective_of_card_eq_natDegree [IsDomain A] [IsDomain B] {p {f : A →+* B} (hf : Function.Injective f) (hroots : Multiset.card p.roots = p.natDegree) : p.roots.map f = (p.map f).roots := by apply Multiset.eq_of_le_of_card_le (map_roots_le_of_injective p hf) - simpa only [Multiset.card_map, hroots] using (card_roots' _).trans (natDegree_map_le f p) + simpa only [Multiset.card_map, hroots] using (card_roots' _).trans natDegree_map_le theorem roots_map_of_map_ne_zero_of_card_eq_natDegree [IsDomain A] [IsDomain B] {p : A[X]} (f : A →+* B) (h : p.map f ≠ 0) (hroots : p.roots.card = p.natDegree) : p.roots.map f = (p.map f).roots := eq_of_le_of_card_le (map_roots_le h) <| by - simpa only [Multiset.card_map, hroots] using (p.map f).card_roots'.trans (p.natDegree_map_le f) + simpa only [Multiset.card_map, hroots] using (p.map f).card_roots'.trans natDegree_map_le theorem Monic.roots_map_of_card_eq_natDegree [IsDomain A] [IsDomain B] {p : A[X]} (hm : p.Monic) (f : A →+* B) (hroots : p.roots.card = p.natDegree) : p.roots.map f = (p.map f).roots := diff --git a/Mathlib/Algebra/Polynomial/Splits.lean b/Mathlib/Algebra/Polynomial/Splits.lean index 211e8294de371..19552f5e13bb4 100644 --- a/Mathlib/Algebra/Polynomial/Splits.lean +++ b/Mathlib/Algebra/Polynomial/Splits.lean @@ -77,7 +77,7 @@ theorem splits_of_degree_le_one {f : K[X]} (hf : degree f ≤ 1) : Splits i f := else by push_neg at hif rw [← Order.succ_le_iff, ← WithBot.coe_zero, WithBot.succ_coe, Nat.succ_eq_succ] at hif - exact splits_of_map_degree_eq_one i (le_antisymm ((degree_map_le i _).trans hf) hif) + exact splits_of_map_degree_eq_one i ((degree_map_le.trans hf).antisymm hif) theorem splits_of_degree_eq_one {f : K[X]} (hf : degree f = 1) : Splits i f := splits_of_degree_le_one i hf.le @@ -189,7 +189,7 @@ but its conditions are easier to check. -/ theorem Splits.comp_of_degree_le_one {f : K[X]} {p : K[X]} (hd : p.degree ≤ 1) (h : f.Splits i) : (f.comp p).Splits i := - Splits.comp_of_map_degree_le_one ((degree_map_le i _).trans hd) h + Splits.comp_of_map_degree_le_one (degree_map_le.trans hd) h theorem Splits.comp_X_sub_C (a : K) {f : K[X]} (h : f.Splits i) : (f.comp (X - C a)).Splits i := diff --git a/Mathlib/Algebra/Polynomial/SumIteratedDerivative.lean b/Mathlib/Algebra/Polynomial/SumIteratedDerivative.lean index f05a9ab0735d1..890f08beece92 100644 --- a/Mathlib/Algebra/Polynomial/SumIteratedDerivative.lean +++ b/Mathlib/Algebra/Polynomial/SumIteratedDerivative.lean @@ -217,14 +217,17 @@ theorem aeval_sumIDeriv_of_pos [Nontrivial A] [NoZeroDivisors A] (p : R[X]) {q : intro r p' hp have : range (p.natDegree + 1) = range q ∪ Ico q (p.natDegree + 1) := by rw [range_eq_Ico, Ico_union_Ico_eq_Ico hq.le] - have h := natDegree_map_le (algebraMap R A) p - rw [congr_arg natDegree hp, natDegree_mul, natDegree_pow, natDegree_X_sub_C, mul_one, - ← Nat.sub_add_comm (Nat.one_le_of_lt hq), tsub_le_iff_right] at h - exact le_of_add_le_left h - · exact pow_ne_zero _ (X_sub_C_ne_zero r) - · rintro rfl - rw [mul_zero, Polynomial.map_eq_zero_iff inj_amap] at hp - exact p0 hp + rw [← tsub_le_iff_right] + calc + q - 1 ≤ q - 1 + p'.natDegree := le_self_add + _ = (p.map <| algebraMap R A).natDegree := by + rw [hp, natDegree_mul, natDegree_pow, natDegree_X_sub_C, mul_one, + ← Nat.sub_add_comm (Nat.one_le_of_lt hq)] + · exact pow_ne_zero _ (X_sub_C_ne_zero r) + · rintro rfl + rw [mul_zero, Polynomial.map_eq_zero_iff inj_amap] at hp + exact p0 hp + _ ≤ p.natDegree := natDegree_map_le rw [← zero_add ((q - 1)! • p'.eval r)] rw [sumIDeriv_apply, map_sum, map_sum, this] have : range q = range (q - 1 + 1) := by rw [tsub_add_cancel_of_le (Nat.one_le_of_lt hq)] diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index bb1699add44f8..2a631d1f70aeb 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -880,7 +880,7 @@ lemma isIntegral_of_isClosedMap_comap_mapRingHom (h : IsClosedMap (comap (mapRin · have : p.natDegree ≤ 1 := by simpa using natDegree_linear_le (a := r) (b := -1) rw [eval₂_eq_eval_map, reverse, Polynomial.map_mul, ← reflect_map, Polynomial.map_pow, map_X, ← revAt_zero (1 + _), ← reflect_monomial, - ← reflect_mul _ _ (natDegree_map_le _ _) (by simp), pow_zero, mul_one, hc, + ← reflect_mul _ _ natDegree_map_le (by simp), pow_zero, mul_one, hc, ← add_assoc, reflect_mul _ _ (this.trans (by simp)) le_rfl, eval_mul, reflect_sub, reflect_mul _ _ (by simp) (by simp)] simp [← pow_succ'] diff --git a/Mathlib/FieldTheory/Minpoly/MinpolyDiv.lean b/Mathlib/FieldTheory/Minpoly/MinpolyDiv.lean index fae1bec1703c0..cdeabc8d727bf 100644 --- a/Mathlib/FieldTheory/Minpoly/MinpolyDiv.lean +++ b/Mathlib/FieldTheory/Minpoly/MinpolyDiv.lean @@ -214,7 +214,7 @@ lemma sum_smul_minpolyDiv_eq_X_pow (E) [Field E] [Algebra K E] [IsAlgClosed E] Finset.mem_univ, forall_true_left, true_and, Finset.fold_max_lt, AlgHom.card] refine ⟨finrank_pos, ?_⟩ intro σ - exact ((Polynomial.natDegree_smul_le _ _).trans (natDegree_map_le _ _)).trans_lt + exact ((Polynomial.natDegree_smul_le _ _).trans natDegree_map_le).trans_lt ((natDegree_minpolyDiv_lt (Algebra.IsIntegral.isIntegral x)).trans_le (minpoly.natDegree_le _)) · rwa [natDegree_pow, natDegree_X, mul_one, AlgHom.card] diff --git a/Mathlib/RingTheory/EisensteinCriterion.lean b/Mathlib/RingTheory/EisensteinCriterion.lean index 0e246b9233e41..775daf2ac12dd 100644 --- a/Mathlib/RingTheory/EisensteinCriterion.lean +++ b/Mathlib/RingTheory/EisensteinCriterion.lean @@ -40,7 +40,7 @@ theorem map_eq_C_mul_X_pow_of_forall_coeff_mem {f : R[X]} {P : Ideal R} · rw [coeff_eq_zero_of_degree_lt, coeff_eq_zero_of_degree_lt] · refine lt_of_le_of_lt (degree_C_mul_X_pow_le _ _) ?_ rwa [← degree_eq_natDegree hf0] - · exact lt_of_le_of_lt (degree_map_le _ _) h + · exact lt_of_le_of_lt degree_map_le h theorem le_natDegree_of_map_eq_mul_X_pow {n : ℕ} {P : Ideal R} (hP : P.IsPrime) {q : R[X]} {c : Polynomial (R ⧸ P)} (hq : map (mk P) q = c * X ^ n) (hc0 : c.degree = 0) : @@ -49,7 +49,7 @@ theorem le_natDegree_of_map_eq_mul_X_pow {n : ℕ} {P : Ideal R} (hP : P.IsPrime (calc ↑n = degree (q.map (mk P)) := by rw [hq, degree_mul, hc0, zero_add, degree_pow, degree_X, nsmul_one] - _ ≤ degree q := degree_map_le _ _ + _ ≤ degree q := degree_map_le _ ≤ natDegree q := degree_le_natDegree ) diff --git a/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean b/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean index 64a1bbcd81c41..99cf15144c642 100644 --- a/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean +++ b/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean @@ -372,7 +372,7 @@ lemma exists_eval_ne_zero_of_totalDegree_le_card_aux {N : ℕ} {F : MvPolynomial have hφR : φ.natDegree < #R := by refine lt_of_lt_of_le ?_ hnR norm_cast - refine lt_of_le_of_lt (natDegree_map_le _ _) ?_ + refine lt_of_le_of_lt natDegree_map_le ?_ suffices (finSuccEquiv _ _ F).natDegree ≠ n by omega rintro rfl refine leadingCoeff_ne_zero.mpr ?_ hFn diff --git a/Mathlib/RingTheory/Polynomial/Eisenstein/Basic.lean b/Mathlib/RingTheory/Polynomial/Eisenstein/Basic.lean index 924247aa7404b..5e2bf7367a632 100644 --- a/Mathlib/RingTheory/Polynomial/Eisenstein/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Eisenstein/Basic.lean @@ -62,7 +62,7 @@ theorem map (hf : f.IsWeaklyEisensteinAt 𝓟) {A : Type v} [CommRing A] (φ : R (f.map φ).IsWeaklyEisensteinAt (𝓟.map φ) := by refine (isWeaklyEisensteinAt_iff _ _).2 fun hn => ?_ rw [coeff_map] - exact mem_map_of_mem _ (hf.mem (lt_of_lt_of_le hn (natDegree_map_le _ _))) + exact mem_map_of_mem _ (hf.mem (lt_of_lt_of_le hn natDegree_map_le)) end CommSemiring @@ -91,7 +91,7 @@ theorem exists_mem_adjoin_mul_eq_pow_natDegree {x : S} (hx : aeval x f = 0) (hmo congr · skip ext i - rw [coeff_map, hφ i.1 (lt_of_lt_of_le i.2 (natDegree_map_le _ _)), + rw [coeff_map, hφ i.1 (lt_of_lt_of_le i.2 natDegree_map_le), RingHom.map_mul, mul_assoc] rw [hx, ← mul_sum, neg_eq_neg_one_mul, ← mul_assoc (-1 : S), mul_comm (-1 : S), mul_assoc] refine diff --git a/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean b/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean index 34a8597d9b708..df43e22b8bed7 100644 --- a/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean +++ b/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean @@ -217,7 +217,7 @@ theorem totient_le_degree_minpoly : Nat.totient n ≤ (minpoly ℤ μ).natDegree _ ≤ P_K.roots.toFinset.card := Finset.card_le_card (is_roots_of_minpoly h) _ ≤ Multiset.card P_K.roots := Multiset.toFinset_card_le _ _ ≤ P_K.natDegree := card_roots' _ - _ ≤ P.natDegree := natDegree_map_le _ _ + _ ≤ P.natDegree := natDegree_map_le end IsDomain From 4bbdccd9c5f862bf90ff12f0a9e2c8be032b9a84 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 2 Dec 2024 01:14:15 +0000 Subject: [PATCH 453/829] chore: bump toolchain to v4.14.0 (#19673) --- lake-manifest.json | 16 ++++++++-------- lakefile.lean | 8 ++++---- lean-toolchain | 2 +- 3 files changed, 13 insertions(+), 13 deletions(-) diff --git a/lake-manifest.json b/lake-manifest.json index 7e9d6679c79f6..81bcd7da5e4e8 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -25,30 +25,30 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "119b022b3ea88ec810a677888528e50f8144a26e", + "rev": "519e509a28864af5bed98033dd33b95cf08e9aa7", "name": "importGraph", "manifestFile": "lake-manifest.json", - "inputRev": "main", + "inputRev": "v4.14.0", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover-community/ProofWidgets4", "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "1383e72b40dd62a566896a6e348ffe868801b172", + "rev": "68280daef58803f68368eb2e53046dabcd270c9d", "name": "proofwidgets", "manifestFile": "lake-manifest.json", - "inputRev": "v0.0.46", + "inputRev": "v0.0.47", "inherited": false, "configFile": "lakefile.lean"}, {"url": "https://github.com/leanprover-community/aesop", "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "de91b59101763419997026c35a41432ac8691f15", + "rev": "5a0ec8588855265ade536f35bcdcf0fb24fd6030", "name": "aesop", "manifestFile": "lake-manifest.json", - "inputRev": "master", + "inputRev": "v4.14.0", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover-community/quote4", @@ -65,10 +65,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "f46c0445cc86ad2bc973edf8c5b2bd88bd91913b", + "rev": "8d6c853f11a5172efa0e96b9f2be1a83d861cdd9", "name": "batteries", "manifestFile": "lake-manifest.json", - "inputRev": "main", + "inputRev": "v4.14.0", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover/lean4-cli", diff --git a/lakefile.lean b/lakefile.lean index cab011aba5693..4f30670824859 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -7,11 +7,11 @@ open Lake DSL ## Mathlib dependencies on upstream projects -/ -require "leanprover-community" / "batteries" @ git "main" +require "leanprover-community" / "batteries" @ git "v4.14.0" require "leanprover-community" / "Qq" @ git "master" -require "leanprover-community" / "aesop" @ git "master" -require "leanprover-community" / "proofwidgets" @ git "v0.0.46" -require "leanprover-community" / "importGraph" @ git "main" +require "leanprover-community" / "aesop" @ git "v4.14.0" +require "leanprover-community" / "proofwidgets" @ git "v0.0.47" +require "leanprover-community" / "importGraph" @ git "v4.14.0" require "leanprover-community" / "LeanSearchClient" @ git "main" from git "https://github.com/leanprover-community/LeanSearchClient" @ "main" require "leanprover-community" / "plausible" @ git "main" diff --git a/lean-toolchain b/lean-toolchain index 6d9e70f73356b..401bc146f623c 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.14.0-rc3 +leanprover/lean4:v4.14.0 \ No newline at end of file From 7c5e371210bc4eb6b54e9771b53d61092b6c3a8a Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Mon, 2 Dec 2024 01:23:53 +0000 Subject: [PATCH 454/829] chore: update Mathlib dependencies 2024-12-02 (#19672) This PR updates the Mathlib dependencies. From dde4f2ecaec222a1c7db638d7536f0b195d8e953 Mon Sep 17 00:00:00 2001 From: hehepig166 <1260647889@qq.com> Date: Mon, 2 Dec 2024 01:32:35 +0000 Subject: [PATCH 455/829] feat: an equality condition for AM-GM inequality (#19435) Add some theorems on the equality condition of the weighted AM-GM inequality for real-valued nonnegative functions, referring to `geom_mean_arith_mean_weighted_eq_iff` in `Mathlib/Analysis/MeanInequalities.lean`. Co-authored-by: hehepig166 <76951962+hehepig166@users.noreply.github.com> --- Mathlib/Analysis/MeanInequalities.lean | 82 +++++++++++++++++++++++++- 1 file changed, 79 insertions(+), 3 deletions(-) diff --git a/Mathlib/Analysis/MeanInequalities.lean b/Mathlib/Analysis/MeanInequalities.lean index 35b6c3d8d8508..ae574c56e55c6 100644 --- a/Mathlib/Analysis/MeanInequalities.lean +++ b/Mathlib/Analysis/MeanInequalities.lean @@ -135,7 +135,7 @@ theorem geom_mean_le_arith_mean_weighted (w z : ι → ℝ) (hw : ∀ i ∈ s, 0 -- for `exp` and numbers `log (z i)` with weights `w i`. · simp only [not_exists, not_and, Ne, Classical.not_not] at A have := convexOn_exp.map_sum_le hw hw' fun i _ => Set.mem_univ <| log (z i) - simp only [exp_sum, (· ∘ ·), smul_eq_mul, mul_comm (w _) (log _)] at this + simp only [exp_sum, smul_eq_mul, mul_comm (w _) (log _)] at this convert this using 1 <;> [apply prod_congr rfl;apply sum_congr rfl] <;> intro i hi · cases' eq_or_lt_of_le (hz i hi) with hz hz · simp [A i hi hz.symm] @@ -144,7 +144,7 @@ theorem geom_mean_le_arith_mean_weighted (w z : ι → ℝ) (hw : ∀ i ∈ s, 0 · simp [A i hi hz.symm] · rw [exp_log hz] -/-- **AM-GM inequality**: The **geometric mean is less than or equal to the arithmetic mean. --/ +/-- **AM-GM inequality**: The **geometric mean is less than or equal to the arithmetic mean. -/ theorem geom_mean_le_arith_mean {ι : Type*} (s : Finset ι) (w : ι → ℝ) (z : ι → ℝ) (hw : ∀ i ∈ s, 0 ≤ w i) (hw' : 0 < ∑ i ∈ s, w i) (hz : ∀ i ∈ s, 0 ≤ z i) : (∏ i ∈ s, z i ^ w i) ^ (∑ i ∈ s, w i)⁻¹ ≤ (∑ i ∈ s, w i * z i) / (∑ i ∈ s, w i) := by @@ -190,6 +190,82 @@ theorem geom_mean_eq_arith_mean_weighted_of_constant (w z : ι → ℝ) (x : ℝ ∏ i ∈ s, z i ^ w i = ∑ i ∈ s, w i * z i := by rw [geom_mean_weighted_of_constant, arith_mean_weighted_of_constant] <;> assumption +/-- **AM-GM inequality - equality condition**: This theorem provides the equality condition for the +*positive* weighted version of the AM-GM inequality for real-valued nonnegative functions. -/ +theorem geom_mean_eq_arith_mean_weighted_iff' (w z : ι → ℝ) (hw : ∀ i ∈ s, 0 < w i) + (hw' : ∑ i ∈ s, w i = 1) (hz : ∀ i ∈ s, 0 ≤ z i) : + ∏ i ∈ s, z i ^ w i = ∑ i ∈ s, w i * z i ↔ ∀ j ∈ s, z j = ∑ i ∈ s, w i * z i := by + by_cases A : ∃ i ∈ s, z i = 0 ∧ w i ≠ 0 + · rcases A with ⟨i, his, hzi, hwi⟩ + rw [prod_eq_zero his] + · constructor + · intro h + rw [← h] + intro j hj + apply eq_zero_of_ne_zero_of_mul_left_eq_zero (ne_of_lt (hw j hj)).symm + apply (sum_eq_zero_iff_of_nonneg ?_).mp h.symm j hj + exact fun i hi => (mul_nonneg_iff_of_pos_left (hw i hi)).mpr (hz i hi) + · intro h + convert h i his + exact hzi.symm + · rw [hzi] + exact zero_rpow hwi + · simp only [not_exists, not_and] at A + have hz' := fun i h => lt_of_le_of_ne (hz i h) (fun a => (A i h a.symm) (ne_of_gt (hw i h))) + have := strictConvexOn_exp.map_sum_eq_iff hw hw' fun i _ => Set.mem_univ <| log (z i) + simp only [exp_sum, smul_eq_mul, mul_comm (w _) (log _)] at this + convert this using 1 + · apply Eq.congr <;> + [apply prod_congr rfl; apply sum_congr rfl] <;> + intro i hi <;> + simp only [exp_mul, exp_log (hz' i hi)] + · constructor <;> intro h j hj + · rw [← arith_mean_weighted_of_constant s w _ (log (z j)) hw' fun i _ => congrFun rfl] + apply sum_congr rfl + intro x hx + simp only [mul_comm, h j hj, h x hx] + · rw [← arith_mean_weighted_of_constant s w _ (z j) hw' fun i _ => congrFun rfl] + apply sum_congr rfl + intro x hx + simp only [log_injOn_pos (hz' j hj) (hz' x hx), h j hj, h x hx] + +/-- **AM-GM inequality - equality condition**: This theorem provides the equality condition for the +weighted version of the AM-GM inequality for real-valued nonnegative functions. -/ +theorem geom_mean_eq_arith_mean_weighted_iff (w z : ι → ℝ) (hw : ∀ i ∈ s, 0 ≤ w i) + (hw' : ∑ i ∈ s, w i = 1) (hz : ∀ i ∈ s, 0 ≤ z i) : + ∏ i ∈ s, z i ^ w i = ∑ i ∈ s, w i * z i ↔ ∀ j ∈ s, w j ≠ 0 → z j = ∑ i ∈ s, w i * z i := by + have h (i) (_ : i ∈ s) : w i * z i ≠ 0 → w i ≠ 0 := by apply left_ne_zero_of_mul + have h' (i) (_ : i ∈ s) : z i ^ w i ≠ 1 → w i ≠ 0 := by + by_contra! + obtain ⟨h1, h2⟩ := this + simp only [h2, rpow_zero, ne_self_iff_false] at h1 + rw [← sum_filter_of_ne h, ← prod_filter_of_ne h', geom_mean_eq_arith_mean_weighted_iff'] + · simp + · simp +contextual [(hw _ _).gt_iff_ne] + · rwa [sum_filter_ne_zero] + · simp_all only [ne_eq, mul_eq_zero, not_or, not_false_eq_true, and_imp, implies_true, mem_filter] + +/-- **AM-GM inequality - strict inequality condition**: This theorem provides the strict inequality +condition for the *positive* weighted version of the AM-GM inequality for real-valued nonnegative +functions. -/ +theorem geom_mean_lt_arith_mean_weighted_iff_of_pos (w z : ι → ℝ) (hw : ∀ i ∈ s, 0 < w i) + (hw' : ∑ i ∈ s, w i = 1) (hz : ∀ i ∈ s, 0 ≤ z i) : + ∏ i ∈ s, z i ^ w i < ∑ i ∈ s, w i * z i ↔ ∃ j ∈ s, ∃ k ∈ s, z j ≠ z k:= by + constructor + · intro h + by_contra! h_contra + rw [(geom_mean_eq_arith_mean_weighted_iff' s w z hw hw' hz).mpr ?_] at h + · exact (lt_self_iff_false _).mp h + · intro j hjs + rw [← arith_mean_weighted_of_constant s w (fun _ => z j) (z j) hw' fun _ _ => congrFun rfl] + apply sum_congr rfl (fun x a => congrArg (HMul.hMul (w x)) (h_contra j hjs x a)) + · rintro ⟨j, hjs, k, hks, hzjk⟩ + have := geom_mean_le_arith_mean_weighted s w z (fun i a => le_of_lt (hw i a)) hw' hz + by_contra! h + apply le_antisymm this at h + apply (geom_mean_eq_arith_mean_weighted_iff' s w z hw hw' hz).mp at h + simp only [h j hjs, h k hks, ne_eq, not_true_eq_false] at hzjk + end Real namespace NNReal @@ -282,7 +358,7 @@ theorem harm_mean_le_geom_mean_weighted (w z : ι → ℝ) (hs : s.Nonempty) (hw · rw [Real.inv_rpow]; apply fun i hi ↦ le_of_lt (hz i hi); assumption -/-- **HM-GM inequality**: The **harmonic mean is less than or equal to the geometric mean. --/ +/-- **HM-GM inequality**: The **harmonic mean is less than or equal to the geometric mean. -/ theorem harm_mean_le_geom_mean {ι : Type*} (s : Finset ι) (hs : s.Nonempty) (w : ι → ℝ) (z : ι → ℝ) (hw : ∀ i ∈ s, 0 < w i) (hw' : 0 < ∑ i in s, w i) (hz : ∀ i ∈ s, 0 < z i) : (∑ i in s, w i) / (∑ i in s, w i / z i) ≤ (∏ i in s, z i ^ w i) ^ (∑ i in s, w i)⁻¹ := by From 41ff1f7899c971f91362710d4444e338b8acd644 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 2 Dec 2024 04:16:57 +0000 Subject: [PATCH 456/829] chore: bump toolchain to v4.15.0-rc1, and merge bump/v4.15.0 adaptations branch (#19675) Co-authored-by: Johan Commelin Co-authored-by: leanprover-community-mathlib4-bot --- Archive/Imo/Imo2024Q5.lean | 14 +- Archive/MiuLanguage/Basic.lean | 4 +- Counterexamples/HomogeneousPrimeNotPrime.lean | 9 +- Mathlib/Algebra/BigOperators/Group/List.lean | 11 +- .../Algebra/Category/Ring/Constructions.lean | 4 +- Mathlib/Algebra/Category/Ring/Epi.lean | 5 +- Mathlib/Algebra/CharP/IntermediateField.lean | 1 - .../Algebra/Group/Pointwise/Finset/Basic.lean | 4 +- Mathlib/Algebra/Group/Subgroup/Ker.lean | 8 +- .../Algebra/Group/Submonoid/Operations.lean | 8 +- Mathlib/Algebra/Homology/Embedding/Basic.lean | 2 +- Mathlib/Algebra/IsPrimePow.lean | 2 +- Mathlib/Algebra/Lie/Semisimple/Basic.lean | 6 +- Mathlib/Algebra/LinearRecurrence.lean | 1 - Mathlib/Algebra/Module/FreeLocus.lean | 4 +- .../Algebra/Order/Group/Unbundled/Basic.lean | 1 + Mathlib/Algebra/Quaternion.lean | 4 +- Mathlib/Algebra/Ring/Subring/Defs.lean | 7 +- Mathlib/Algebra/Tropical/Basic.lean | 4 +- .../Morphisms/QuasiCompact.lean | 1 - .../Morphisms/QuasiSeparated.lean | 1 - Mathlib/AlgebraicGeometry/Spec.lean | 6 +- .../ContinuousFunctionalCalculus/Order.lean | 5 +- .../Calculus/ContDiff/FaaDiBruno.lean | 2 +- Mathlib/Analysis/Calculus/FDeriv/Mul.lean | 4 +- .../Analysis/Calculus/FDeriv/Symmetric.lean | 2 + .../Complex/UpperHalfPlane/Metric.lean | 10 +- .../FunctionalSpaces/SobolevInequality.lean | 10 +- .../Normed/Operator/BoundedLinearMaps.lean | 11 +- .../NormedSpace/Multilinear/Basic.lean | 2 +- Mathlib/Analysis/SumOverResidueClass.lean | 2 +- Mathlib/CategoryTheory/GlueData.lean | 2 +- Mathlib/CategoryTheory/Monoidal/Mod_.lean | 8 ++ Mathlib/CategoryTheory/Monoidal/Mon_.lean | 8 ++ Mathlib/CategoryTheory/Sites/Adjunction.lean | 4 +- .../Additive/PluenneckeRuzsa.lean | 8 +- .../Combinatorics/Additive/RuzsaCovering.lean | 12 +- Mathlib/Combinatorics/Colex.lean | 2 +- .../Enumerative/Composition.lean | 4 +- Mathlib/Computability/Ackermann.lean | 2 +- Mathlib/Computability/Halting.lean | 28 ++-- Mathlib/Computability/Partrec.lean | 21 +-- Mathlib/Computability/PartrecCode.lean | 2 +- Mathlib/Computability/Primrec.lean | 38 +++--- Mathlib/Computability/TMToPartrec.lean | 22 +-- Mathlib/Computability/TuringMachine.lean | 11 +- Mathlib/Control/Random.lean | 2 +- Mathlib/Data/Bool/Basic.lean | 4 +- Mathlib/Data/DFinsupp/Sigma.lean | 10 +- Mathlib/Data/ENNReal/Inv.lean | 2 +- Mathlib/Data/Fin/Tuple/Take.lean | 1 - Mathlib/Data/Finset/Defs.lean | 2 + Mathlib/Data/Finset/SDiff.lean | 6 +- Mathlib/Data/Finset/Sort.lean | 4 +- Mathlib/Data/Fintype/BigOperators.lean | 3 +- Mathlib/Data/Fintype/Fin.lean | 2 +- Mathlib/Data/Fintype/Vector.lean | 2 +- Mathlib/Data/List/Basic.lean | 38 ++---- Mathlib/Data/List/Cycle.lean | 10 +- Mathlib/Data/List/FinRange.lean | 1 - Mathlib/Data/List/GetD.lean | 2 +- Mathlib/Data/List/Indexes.lean | 12 +- Mathlib/Data/List/InsertIdx.lean | 128 +----------------- Mathlib/Data/List/Intervals.lean | 2 +- Mathlib/Data/List/NodupEquivFin.lean | 8 +- Mathlib/Data/List/OfFn.lean | 53 +------- Mathlib/Data/List/Permutation.lean | 18 +-- Mathlib/Data/List/Sublists.lean | 6 +- Mathlib/Data/List/Sym.lean | 2 +- Mathlib/Data/List/TFAE.lean | 2 +- Mathlib/Data/Multiset/Basic.lean | 2 +- Mathlib/Data/NNReal/Defs.lean | 2 +- Mathlib/Data/Nat/BitIndices.lean | 2 +- Mathlib/Data/Nat/Bitwise.lean | 2 +- Mathlib/Data/Nat/Choose/Lucas.lean | 4 +- Mathlib/Data/Nat/Count.lean | 1 + Mathlib/Data/Nat/Defs.lean | 15 +- Mathlib/Data/Nat/Factorization/Basic.lean | 4 +- Mathlib/Data/Nat/Log.lean | 2 +- Mathlib/Data/Nat/Prime/Basic.lean | 9 +- Mathlib/Data/Num/Bitwise.lean | 2 +- Mathlib/Data/Rat/Denumerable.lean | 4 +- Mathlib/Data/Seq/Seq.lean | 2 +- Mathlib/Data/Set/List.lean | 6 +- Mathlib/Data/Sign.lean | 8 +- Mathlib/Data/Sym/Basic.lean | 15 +- Mathlib/Data/Sym/Sym2.lean | 2 +- Mathlib/Data/Vector/Basic.lean | 4 +- Mathlib/Data/Vector/Mem.lean | 4 +- Mathlib/Deprecated/Order.lean | 2 +- Mathlib/Dynamics/Newton.lean | 2 +- Mathlib/FieldTheory/IsPerfectClosure.lean | 2 +- Mathlib/FieldTheory/KummerExtension.lean | 2 +- Mathlib/FieldTheory/PerfectClosure.lean | 2 +- Mathlib/FieldTheory/Relrank.lean | 11 ++ Mathlib/FieldTheory/SeparableDegree.lean | 2 +- Mathlib/Geometry/Euclidean/MongePoint.lean | 34 ++--- Mathlib/Geometry/Manifold/Instances/Real.lean | 3 + .../Manifold/MFDeriv/NormedSpace.lean | 40 +++++- .../RingedSpace/PresheafedSpace/Gluing.lean | 2 +- Mathlib/GroupTheory/Congruence/Basic.lean | 2 +- Mathlib/GroupTheory/Coxeter/Inversion.lean | 4 +- Mathlib/GroupTheory/Perm/Cycle/Type.lean | 10 +- Mathlib/Lean/Meta/DiscrTree.lean | 12 +- Mathlib/Lean/Meta/KAbstractPositions.lean | 2 +- Mathlib/Lean/Meta/Simp.lean | 8 +- Mathlib/LinearAlgebra/Prod.lean | 3 +- .../LinearAlgebra/TensorProduct/Basic.lean | 8 +- Mathlib/Logic/Equiv/List.lean | 4 +- Mathlib/Logic/Small/List.lean | 4 +- Mathlib/MeasureTheory/Function/LpSpace.lean | 3 - Mathlib/MeasureTheory/Group/Measure.lean | 7 +- .../MeasurableSpace/Embedding.lean | 11 +- .../Measure/Lebesgue/VolumeOfBalls.lean | 4 +- Mathlib/ModelTheory/Encoding.lean | 2 +- Mathlib/ModelTheory/Syntax.lean | 1 + .../ClassNumber/AdmissibleAbsoluteValue.lean | 2 +- .../NumberTheory/FactorisationProperties.lean | 2 +- Mathlib/NumberTheory/FermatPsp.lean | 2 +- .../NumberTheory/LSeries/Nonvanishing.lean | 9 +- .../NumberTheory/Padics/PadicIntegers.lean | 2 +- Mathlib/NumberTheory/SmoothNumbers.lean | 6 +- Mathlib/Order/Filter/Defs.lean | 2 +- Mathlib/Order/Interval/Finset/Box.lean | 5 - Mathlib/Order/WellFounded.lean | 12 +- Mathlib/RingTheory/Algebraic/Basic.lean | 2 +- Mathlib/RingTheory/AlgebraicIndependent.lean | 4 +- Mathlib/RingTheory/Ideal/Operations.lean | 16 ++- Mathlib/RingTheory/Jacobson.lean | 7 +- Mathlib/RingTheory/Multiplicity.lean | 2 +- .../Polynomial/UniqueFactorization.lean | 2 +- Mathlib/RingTheory/WittVector/Frobenius.lean | 2 +- Mathlib/RingTheory/WittVector/InitTail.lean | 2 +- Mathlib/SetTheory/Cardinal/Basic.lean | 7 +- Mathlib/SetTheory/Game/PGame.lean | 1 - Mathlib/SetTheory/Ordinal/Basic.lean | 11 +- Mathlib/Tactic/Abel.lean | 8 +- Mathlib/Tactic/CC/Addition.lean | 4 +- Mathlib/Tactic/CancelDenoms/Core.lean | 2 +- .../CategoryTheory/BicategoryCoherence.lean | 2 +- Mathlib/Tactic/CategoryTheory/Coherence.lean | 2 +- .../Tactic/CategoryTheory/Elementwise.lean | 2 +- Mathlib/Tactic/DeprecateTo.lean | 2 +- Mathlib/Tactic/DeriveTraversable.lean | 10 +- Mathlib/Tactic/FBinop.lean | 4 +- Mathlib/Tactic/FieldSimp.lean | 50 +++---- Mathlib/Tactic/FunProp/Attr.lean | 1 - Mathlib/Tactic/FunProp/Core.lean | 12 +- Mathlib/Tactic/FunProp/Decl.lean | 4 +- Mathlib/Tactic/FunProp/FunctionData.lean | 5 +- Mathlib/Tactic/FunProp/Mor.lean | 14 +- Mathlib/Tactic/FunProp/RefinedDiscrTree.lean | 62 ++++----- Mathlib/Tactic/FunProp/Theorems.lean | 40 ++++-- Mathlib/Tactic/FunProp/ToBatteries.lean | 2 +- Mathlib/Tactic/Linarith/Frontend.lean | 2 +- .../Oracle/SimplexAlgorithm/Datatypes.lean | 4 +- .../SimplexAlgorithm/SimplexAlgorithm.lean | 6 +- Mathlib/Tactic/Linarith/Preprocessing.lean | 2 +- Mathlib/Tactic/LinearCombination'.lean | 2 +- Mathlib/Tactic/LinearCombination.lean | 2 +- Mathlib/Tactic/Linter/PPRoundtrip.lean | 2 +- Mathlib/Tactic/MinImports.lean | 4 +- Mathlib/Tactic/Module.lean | 5 +- Mathlib/Tactic/MoveAdd.lean | 5 +- Mathlib/Tactic/NormNum/Basic.lean | 2 +- Mathlib/Tactic/NormNum/BigOperators.lean | 2 +- Mathlib/Tactic/NormNum/Core.lean | 16 +-- Mathlib/Tactic/NormNum/DivMod.lean | 2 +- Mathlib/Tactic/NormNum/OfScientific.lean | 2 +- Mathlib/Tactic/NormNum/Pow.lean | 2 +- Mathlib/Tactic/NormNum/PowMod.lean | 2 +- Mathlib/Tactic/NormNum/Result.lean | 2 +- Mathlib/Tactic/Polyrith.lean | 2 +- Mathlib/Tactic/Positivity/Core.lean | 9 +- Mathlib/Tactic/Propose.lean | 7 +- Mathlib/Tactic/PushNeg.lean | 7 +- Mathlib/Tactic/ReduceModChar.lean | 7 +- Mathlib/Tactic/Relation/Rfl.lean | 2 +- Mathlib/Tactic/Relation/Symm.lean | 2 +- Mathlib/Tactic/Ring/RingNF.lean | 11 +- Mathlib/Tactic/SimpIntro.lean | 2 +- Mathlib/Tactic/Simps/Basic.lean | 9 +- Mathlib/Tactic/SplitIfs.lean | 2 +- Mathlib/Tactic/Spread.lean | 2 - Mathlib/Tactic/ToAdditive/Frontend.lean | 2 +- Mathlib/Tactic/Variable.lean | 4 +- Mathlib/Tactic/WLOG.lean | 3 +- Mathlib/Testing/Plausible/Functions.lean | 2 +- Mathlib/Topology/Algebra/PontryaginDual.lean | 2 +- Mathlib/Topology/Gluing.lean | 2 +- Mathlib/Topology/List.lean | 20 +-- .../Topology/OmegaCompletePartialOrder.lean | 2 +- Mathlib/Util/DischargerAsTactic.lean | 2 +- MathlibTest/Expr.lean | 2 +- MathlibTest/NthRewrite.lean | 2 +- MathlibTest/Simps.lean | 2 +- MathlibTest/cc.lean | 12 +- lake-manifest.json | 24 ++-- lakefile.lean | 13 +- lean-toolchain | 2 +- 200 files changed, 698 insertions(+), 772 deletions(-) diff --git a/Archive/Imo/Imo2024Q5.lean b/Archive/Imo/Imo2024Q5.lean index 5cbdf23de02f8..4fa8394b08d3d 100644 --- a/Archive/Imo/Imo2024Q5.lean +++ b/Archive/Imo/Imo2024Q5.lean @@ -229,7 +229,7 @@ lemma find?_eq_eq_find?_le {l : List (Cell N)} {r : Fin (N + 2)} (hne : l ≠ [] by_cases h : head.1 = r · simp [h] · have h' : ¬(r ≤ head.1) := fun hr' ↦ h (le_antisymm hf hr') - simp only [h, decide_False, Bool.false_eq_true, not_false_eq_true, List.find?_cons_of_neg, h'] + simp only [h, decide_false, Bool.false_eq_true, not_false_eq_true, List.find?_cons_of_neg, h'] rcases tail with ⟨⟩ | ⟨htail, ttail⟩ · simp · simp only [List.chain'_cons] at ha @@ -315,7 +315,7 @@ lemma Path.tail_findFstEq (p : Path N) {r : Fin (N + 2)} (hr : r ≠ 0) : rcases cells with ⟨⟩ | ⟨head, tail⟩ · simp at nonempty · simp only [List.head_cons] at head_first_row - simp only [List.find?_cons, head_first_row, hr.symm, decide_False] + simp only [List.find?_cons, head_first_row, hr.symm, decide_false] rfl · simp_rw [Path.tail, if_neg h] @@ -329,7 +329,7 @@ lemma Path.tail_firstMonster (p : Path N) (m : MonsterData N) : · simp at nonempty · simp only [List.head_cons] at head_first_row simp only [List.find?_cons, head_first_row, - m.not_mem_monsterCells_of_fst_eq_zero head_first_row, decide_False] + m.not_mem_monsterCells_of_fst_eq_zero head_first_row, decide_false] rfl · simp_rw [Path.tail, if_neg h] @@ -352,17 +352,17 @@ lemma Path.firstMonster_eq_of_findFstEq_mem {p : Path N} {m : MonsterData N} · simp only [List.head_cons] at head_first_row simp only [List.getElem_cons_succ] at h1 simp only [List.length_cons, lt_add_iff_pos_left, List.length_pos_iff_ne_nil] at hl - simp only [m.not_mem_monsterCells_of_fst_eq_zero head_first_row, decide_False, + simp only [m.not_mem_monsterCells_of_fst_eq_zero head_first_row, decide_false, Bool.false_eq_true, not_false_eq_true, List.find?_cons_of_neg, head_first_row, Fin.zero_eq_one_iff, Nat.reduceEqDiff, Option.some_get] - simp only [findFstEq, head_first_row, Fin.zero_eq_one_iff, Nat.reduceEqDiff, decide_False, + simp only [findFstEq, head_first_row, Fin.zero_eq_one_iff, Nat.reduceEqDiff, decide_false, Bool.false_eq_true, not_false_eq_true, List.find?_cons_of_neg] at h rcases tail with ⟨⟩ | ⟨htail, ttail⟩ · simp at hl · simp only [List.getElem_cons_zero] at h1 have h1' : htail.1 = 1 := by simp [Fin.ext_iff, h1] - simp only [h1', decide_True, List.find?_cons_of_pos, Option.get_some] at h - simp only [h1', h, decide_True, List.find?_cons_of_pos] + simp only [h1', decide_true, List.find?_cons_of_pos, Option.get_some] at h + simp only [h1', h, decide_true, List.find?_cons_of_pos] case ind p ht => have h1 : (1 : Fin (N + 2)) ≠ 0 := by simp rw [p.tail_findFstEq h1, p.tail_firstMonster m] at ht diff --git a/Archive/MiuLanguage/Basic.lean b/Archive/MiuLanguage/Basic.lean index 1b1ac4a10368e..722c0eeebd657 100644 --- a/Archive/MiuLanguage/Basic.lean +++ b/Archive/MiuLanguage/Basic.lean @@ -106,9 +106,7 @@ instance : Repr MiuAtom := /-- For simplicity, an `Miustr` is just a list of elements of type `MiuAtom`. -/ -def Miustr := - List MiuAtom -deriving Append +abbrev Miustr := List MiuAtom instance : Membership MiuAtom Miustr := by unfold Miustr; infer_instance diff --git a/Counterexamples/HomogeneousPrimeNotPrime.lean b/Counterexamples/HomogeneousPrimeNotPrime.lean index 1499232c5f88f..61c133a80e292 100644 --- a/Counterexamples/HomogeneousPrimeNotPrime.lean +++ b/Counterexamples/HomogeneousPrimeNotPrime.lean @@ -110,13 +110,8 @@ theorem grading.right_inv : Function.RightInverse (coeLinearMap (grading R)) gra induction' zz using DirectSum.induction_on with i zz d1 d2 ih1 ih2 · simp only [map_zero] · rcases i with (_ | ⟨⟨⟩⟩) <;> rcases zz with ⟨⟨a, b⟩, hab : _ = _⟩ <;> dsimp at hab <;> - cases hab <;> - -- Porting note: proof was `decide` - -- now we need a `simp` and two `erw` subproofs... - simp only [coeLinearMap_of, decompose, AddMonoidHom.coe_mk, - ZeroHom.coe_mk, sub_self, sub_zero] - · erw [map_zero (of (grading R ·) 1), add_zero]; rfl - · erw [map_zero (of (grading R ·) 0), zero_add]; rfl + -- Porting note: proof was `decide` (without reverting any free variables). + cases hab <;> decide +revert · simp only [map_add, ih1, ih2] theorem grading.left_inv : Function.LeftInverse (coeLinearMap (grading R)) grading.decompose := diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index 2bd6c8a351d66..71307df2dbf18 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -205,7 +205,8 @@ theorem prod_take_mul_prod_drop (L : List M) (i : ℕ) : @[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] := by - simp [← take_concat_get', p] + rw [← take_concat_get' _ _ p, prod_append] + simp /-- A list with product not one must have positive length. -/ @[to_additive "A list with sum not zero must have positive length."] @@ -386,10 +387,10 @@ lemma prod_map_ite (p : α → Prop) [DecidablePred p] (f g : α → M) (l : Lis rw [ih] clear ih by_cases hx : p x - · simp only [hx, ↓reduceIte, decide_not, decide_True, map_cons, prod_cons, not_true_eq_false, - decide_False, Bool.false_eq_true, mul_assoc] - · simp only [hx, ↓reduceIte, decide_not, decide_False, Bool.false_eq_true, not_false_eq_true, - decide_True, map_cons, prod_cons, mul_left_comm] + · simp only [hx, ↓reduceIte, decide_not, decide_true, map_cons, prod_cons, not_true_eq_false, + decide_false, Bool.false_eq_true, mul_assoc] + · simp only [hx, ↓reduceIte, decide_not, decide_false, Bool.false_eq_true, not_false_eq_true, + decide_true, map_cons, prod_cons, mul_left_comm] @[to_additive] lemma prod_map_filter_mul_prod_map_filter_not (p : α → Prop) [DecidablePred p] (f : α → M) diff --git a/Mathlib/Algebra/Category/Ring/Constructions.lean b/Mathlib/Algebra/Category/Ring/Constructions.lean index c124625d6a67f..8d91389010d43 100644 --- a/Mathlib/Algebra/Category/Ring/Constructions.lean +++ b/Mathlib/Algebra/Category/Ring/Constructions.lean @@ -106,8 +106,8 @@ def pushoutCoconeIsColimit : Limits.IsColimit (pushoutCocone R A B) := lemma isPushout_tensorProduct (R A B : Type u) [CommRing R] [CommRing A] [CommRing B] [Algebra R A] [Algebra R B] : IsPushout (ofHom <| algebraMap R A) (ofHom <| algebraMap R B) - (ofHom <| Algebra.TensorProduct.includeLeftRingHom (R := R)) - (ofHom <| Algebra.TensorProduct.includeRight.toRingHom (R := R)) where + (ofHom (S := A ⊗[R] B) <| Algebra.TensorProduct.includeLeftRingHom) + (ofHom (S := A ⊗[R] B) <| Algebra.TensorProduct.includeRight.toRingHom) where w := by ext simp diff --git a/Mathlib/Algebra/Category/Ring/Epi.lean b/Mathlib/Algebra/Category/Ring/Epi.lean index eaf70a8f1b5dc..a961e33cdf72e 100644 --- a/Mathlib/Algebra/Category/Ring/Epi.lean +++ b/Mathlib/Algebra/Category/Ring/Epi.lean @@ -23,8 +23,11 @@ lemma CommRingCat.epi_iff_tmul_eq_tmul {R S : Type u} [CommRing R] [CommRing S] ∀ s : S, s ⊗ₜ[R] 1 = 1 ⊗ₜ s := by constructor · intro H + #adaptation_note + /-- After https://github.com/leanprover/lean4/pull/6024 + we need to add `(R := R) (A := S)` in the next line to deal with unification issues. -/ have := H.1 (CommRingCat.ofHom <| Algebra.TensorProduct.includeLeftRingHom (R := R)) - (CommRingCat.ofHom <| Algebra.TensorProduct.includeRight.toRingHom) + (CommRingCat.ofHom <| (Algebra.TensorProduct.includeRight (R := R) (A := S)).toRingHom) (by ext r; show algebraMap R S r ⊗ₜ 1 = 1 ⊗ₜ algebraMap R S r; simp only [Algebra.algebraMap_eq_smul_one, smul_tmul]) exact RingHom.congr_fun this diff --git a/Mathlib/Algebra/CharP/IntermediateField.lean b/Mathlib/Algebra/CharP/IntermediateField.lean index 611dc38c84f62..6016d298478c6 100644 --- a/Mathlib/Algebra/CharP/IntermediateField.lean +++ b/Mathlib/Algebra/CharP/IntermediateField.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jz Pan -/ import Mathlib.Algebra.CharP.Algebra -import Mathlib.Algebra.EuclideanDomain.Field import Mathlib.FieldTheory.IntermediateField.Basic /-! diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index 082727c5e5d61..2e78e81394ffc 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -958,7 +958,7 @@ protected lemma pow_right_monotone (hs : 1 ∈ s) : Monotone (s ^ ·) := pow_right_monotone <| one_subset.2 hs @[to_additive (attr := gcongr)] -lemma pow_subset_pow_left (hst : s ⊆ t) : s ^ n ⊆ t ^ n := pow_left_mono _ hst +lemma pow_subset_pow_left (hst : s ⊆ t) : s ^ n ⊆ t ^ n := subset_of_le (pow_left_mono n hst) @[to_additive (attr := gcongr)] lemma pow_subset_pow_right (hs : 1 ∈ s) (hmn : m ≤ n) : s ^ m ⊆ s ^ n := @@ -979,7 +979,7 @@ alias nsmul_subset_nsmul_of_zero_mem := nsmul_subset_nsmul_right @[to_additive] lemma pow_subset_pow_mul_of_sq_subset_mul (hst : s ^ 2 ⊆ t * s) (hn : n ≠ 0) : - s ^ n ⊆ t ^ (n - 1) * s := pow_le_pow_mul_of_sq_le_mul hst hn + s ^ n ⊆ t ^ (n - 1) * s := subset_of_le (pow_le_pow_mul_of_sq_le_mul hst hn) @[to_additive (attr := simp) nsmul_empty] lemma empty_pow (hn : n ≠ 0) : (∅ : Finset α) ^ n = ∅ := match n with | n + 1 => by simp [pow_succ] diff --git a/Mathlib/Algebra/Group/Subgroup/Ker.lean b/Mathlib/Algebra/Group/Subgroup/Ker.lean index 5ba75d983b6b7..8fd902758f7c1 100644 --- a/Mathlib/Algebra/Group/Subgroup/Ker.lean +++ b/Mathlib/Algebra/Group/Subgroup/Ker.lean @@ -144,9 +144,15 @@ theorem range_one : (1 : G →* N).range = ⊥ := theorem _root_.Subgroup.range_subtype (H : Subgroup G) : H.subtype.range = H := SetLike.coe_injective <| (coe_range _).trans <| Subtype.range_coe -@[to_additive (attr := deprecated (since := "2024-11-26"))] +@[to_additive] alias _root_.Subgroup.subtype_range := Subgroup.range_subtype +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated Subgroup.range_subtype (since := "2024-11-26")] _root_.Subgroup.subtype_range +attribute [deprecated AddSubgroup.range_subtype (since := "2024-11-26")] +_root_.AddSubgroup.subtype_range + @[to_additive (attr := simp)] theorem _root_.Subgroup.inclusion_range {H K : Subgroup G} (h_le : H ≤ K) : (inclusion h_le).range = H.subgroupOf K := diff --git a/Mathlib/Algebra/Group/Submonoid/Operations.lean b/Mathlib/Algebra/Group/Submonoid/Operations.lean index 86c4066e007df..a643434cac8a2 100644 --- a/Mathlib/Algebra/Group/Submonoid/Operations.lean +++ b/Mathlib/Algebra/Group/Submonoid/Operations.lean @@ -870,7 +870,13 @@ def inclusion {S T : Submonoid M} (h : S ≤ T) : S →* T := theorem mrange_subtype (s : Submonoid M) : mrange s.subtype = s := SetLike.coe_injective <| (coe_mrange _).trans <| Subtype.range_coe -@[to_additive (attr := deprecated (since := "2024-11-25"))] alias range_subtype := mrange_subtype +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +@[to_additive] alias range_subtype := mrange_subtype +attribute [deprecated mrange_subtype (since := "2024-11-25")] range_subtype +attribute [deprecated AddSubmonoid.mrange_subtype (since := "2024-11-25")] +AddSubmonoid.range_subtype + @[to_additive] theorem eq_top_iff' : S = ⊤ ↔ ∀ x : M, x ∈ S := diff --git a/Mathlib/Algebra/Homology/Embedding/Basic.lean b/Mathlib/Algebra/Homology/Embedding/Basic.lean index e36261785b6ea..4de24385f83a8 100644 --- a/Mathlib/Algebra/Homology/Embedding/Basic.lean +++ b/Mathlib/Algebra/Homology/Embedding/Basic.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ import Mathlib.Algebra.Homology.ComplexShape -import Mathlib.Algebra.Order.Ring.Nat import Mathlib.Algebra.Ring.Int.Defs +import Mathlib.Algebra.Ring.Nat /-! # Embeddings of complex shapes diff --git a/Mathlib/Algebra/IsPrimePow.lean b/Mathlib/Algebra/IsPrimePow.lean index c21cf0363fe86..d9de71addc9b7 100644 --- a/Mathlib/Algebra/IsPrimePow.lean +++ b/Mathlib/Algebra/IsPrimePow.lean @@ -71,7 +71,7 @@ theorem isPrimePow_nat_iff_bounded (n : ℕ) : rw [isPrimePow_nat_iff] refine Iff.symm ⟨fun ⟨p, _, k, _, hp, hk, hn⟩ => ⟨p, k, hp, hk, hn⟩, ?_⟩ rintro ⟨p, k, hp, hk, rfl⟩ - refine ⟨p, ?_, k, (Nat.lt_pow_self hp.one_lt _).le, hp, hk, rfl⟩ + refine ⟨p, ?_, k, (Nat.lt_pow_self hp.one_lt).le, hp, hk, rfl⟩ conv => { lhs; rw [← (pow_one p)] } exact Nat.pow_le_pow_right hp.one_lt.le hk diff --git a/Mathlib/Algebra/Lie/Semisimple/Basic.lean b/Mathlib/Algebra/Lie/Semisimple/Basic.lean index 50b8820a1addc..791f324394561 100644 --- a/Mathlib/Algebra/Lie/Semisimple/Basic.lean +++ b/Mathlib/Algebra/Lie/Semisimple/Basic.lean @@ -212,7 +212,11 @@ lemma finitelyAtomistic : ∀ s : Finset (LieIdeal R L), ↑s ⊆ {I : LieIdeal set K := s'.sup id suffices I ≤ K by obtain ⟨t, hts', htI⟩ := finitelyAtomistic s' hs'S I this - exact ⟨t, hts'.trans hs'.subset, htI⟩ + #adaptation_note + /-- Prior to https://github.com/leanprover/lean4/pull/6024 + we could write `hts'.trans hs'.subset` instead of + `Finset.Subset.trans hts' hs'.subset` in the next line. -/ + exact ⟨t, Finset.Subset.trans hts' hs'.subset, htI⟩ -- Since `I` is contained in the supremum of `J` with the supremum of `s'`, -- any element `x` of `I` can be written as `y + z` for some `y ∈ J` and `z ∈ K`. intro x hx diff --git a/Mathlib/Algebra/LinearRecurrence.lean b/Mathlib/Algebra/LinearRecurrence.lean index 86fe25eb54450..696a59e766aa7 100644 --- a/Mathlib/Algebra/LinearRecurrence.lean +++ b/Mathlib/Algebra/LinearRecurrence.lean @@ -96,7 +96,6 @@ theorem eq_mk_of_is_sol_of_eq_init {u : ℕ → α} {init : Fin E.order → α} rw [mkSol] split_ifs with h' · exact mod_cast heq ⟨n, h'⟩ - simp only rw [← tsub_add_cancel_of_le (le_of_not_lt h'), h (n - E.order)] congr with k have : n - E.order + k < n := by diff --git a/Mathlib/Algebra/Module/FreeLocus.lean b/Mathlib/Algebra/Module/FreeLocus.lean index 9f42786132f02..1487895fd48a4 100644 --- a/Mathlib/Algebra/Module/FreeLocus.lean +++ b/Mathlib/Algebra/Module/FreeLocus.lean @@ -126,7 +126,7 @@ lemma freeLocus_localization (S : Submonoid R) : ⟨fun r r' m ↦ show algebraMap _ Rₚ (r • r') • m = _ by simp [Algebra.smul_def, ← IsScalarTower.algebraMap_apply, mul_smul]; rfl⟩ have : IsScalarTower (Localization S) Rₚ Mₚ := - ⟨fun r r' m ↦ show _ = algebraMap _ Rₚ r • _ by rw [← mul_smul, ← Algebra.smul_def]⟩ + ⟨fun r r' m ↦ show _ = algebraMap _ Rₚ r • r' • m by rw [← mul_smul, ← Algebra.smul_def]⟩ let l := (IsLocalizedModule.liftOfLE _ _ hp' (LocalizedModule.mkLinearMap S M) (LocalizedModule.mkLinearMap p'.primeCompl M)).extendScalarsOfIsLocalization S (Localization S) @@ -203,7 +203,7 @@ lemma isLocallyConstant_rankAtStalk_freeLocus [Module.FinitePresentation R M] : ⟨fun r r' m ↦ show algebraMap _ Rₚ (r • r') • m = _ by simp [Algebra.smul_def, ← IsScalarTower.algebraMap_apply, mul_smul]; rfl⟩ have : IsScalarTower (Localization.Away f) Rₚ Mₚ := - ⟨fun r r' m ↦ show _ = algebraMap _ Rₚ r • _ by rw [← mul_smul, ← Algebra.smul_def]⟩ + ⟨fun r r' m ↦ show _ = algebraMap _ Rₚ r • r' • m by rw [← mul_smul, ← Algebra.smul_def]⟩ let l := (IsLocalizedModule.liftOfLE _ _ hp' (LocalizedModule.mkLinearMap (.powers f) M) (LocalizedModule.mkLinearMap p.asIdeal.primeCompl M)).extendScalarsOfIsLocalization (.powers f) (Localization.Away f) diff --git a/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean index 03c9d7579e5f1..5b90caf6e8964 100644 --- a/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean @@ -537,6 +537,7 @@ section LE variable [LE α] [MulLeftMono α] {a b c d : α} +set_option linter.docPrime false in @[to_additive sub_le_sub_iff] theorem div_le_div_iff' : a / b ≤ c / d ↔ a * d ≤ c * b := by simpa only [div_eq_mul_inv] using mul_inv_le_mul_inv_iff' diff --git a/Mathlib/Algebra/Quaternion.lean b/Mathlib/Algebra/Quaternion.lean index 1ee91c3592cc2..e32a58000bb51 100644 --- a/Mathlib/Algebra/Quaternion.lean +++ b/Mathlib/Algebra/Quaternion.lean @@ -744,8 +744,6 @@ namespace Quaternion variable {S T R : Type*} [CommRing R] (r x y : R) (a b : ℍ[R]) -export QuaternionAlgebra (re imI imJ imK) - /-- Coercion `R → ℍ[R]`. -/ @[coe] def coe : R → ℍ[R] := QuaternionAlgebra.coe @@ -1191,7 +1189,7 @@ theorem normSq_le_zero : normSq a ≤ 0 ↔ a = 0 := normSq_nonneg.le_iff_eq.trans normSq_eq_zero instance instNontrivial : Nontrivial ℍ[R] where - exists_pair_ne := ⟨0, 1, mt (congr_arg re) zero_ne_one⟩ + exists_pair_ne := ⟨0, 1, mt (congr_arg QuaternionAlgebra.re) zero_ne_one⟩ instance : NoZeroDivisors ℍ[R] where eq_zero_or_eq_zero_of_mul_eq_zero {a b} hab := diff --git a/Mathlib/Algebra/Ring/Subring/Defs.lean b/Mathlib/Algebra/Ring/Subring/Defs.lean index 653ae172134f9..3b201ab0ecd61 100644 --- a/Mathlib/Algebra/Ring/Subring/Defs.lean +++ b/Mathlib/Algebra/Ring/Subring/Defs.lean @@ -98,15 +98,16 @@ instance (priority := 75) toHasIntCast : IntCast s := -- Prefer subclasses of `Ring` over subclasses of `SubringClass`. /-- A subring of a ring inherits a ring structure -/ instance (priority := 75) toRing : Ring s := - Subtype.coe_injective.ring (↑) rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) + Subtype.coe_injective.ring Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) fun _ => rfl -- Prefer subclasses of `Ring` over subclasses of `SubringClass`. /-- A subring of a `CommRing` is a `CommRing`. -/ instance (priority := 75) toCommRing {R} [CommRing R] [SetLike S R] [SubringClass S R] : CommRing s := - Subtype.coe_injective.commRing (↑) rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) - (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) fun _ => rfl + Subtype.coe_injective.commRing Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) + (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) + (fun _ => rfl) fun _ => rfl -- Prefer subclasses of `Ring` over subclasses of `SubringClass`. /-- A subring of a domain is a domain. -/ diff --git a/Mathlib/Algebra/Tropical/Basic.lean b/Mathlib/Algebra/Tropical/Basic.lean index 5a22f3bbfe005..dfb67530f2fc9 100644 --- a/Mathlib/Algebra/Tropical/Basic.lean +++ b/Mathlib/Algebra/Tropical/Basic.lean @@ -6,8 +6,10 @@ Authors: Yakov Pechersky import Mathlib.Algebra.Order.Monoid.Unbundled.Pow import Mathlib.Algebra.SMulWithZero import Mathlib.Order.Hom.Basic -import Mathlib.Algebra.Order.Ring.Nat import Mathlib.Algebra.Order.Monoid.Unbundled.WithTop +import Mathlib.Algebra.Order.AddGroupWithTop +import Mathlib.Algebra.Ring.Nat +import Mathlib.Algebra.Order.Monoid.Unbundled.MinMax /-! diff --git a/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean b/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean index 18ec910454728..a44c42830da74 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean @@ -116,7 +116,6 @@ theorem isCompact_basicOpen (X : Scheme) {U : X.Opens} (hU : IsCompact (U : Set refine Set.Subset.trans ?_ (Set.subset_iUnion₂ j hj) exact Set.Subset.rfl -@[reducible] instance : HasAffineProperty @QuasiCompact (fun X _ _ _ ↦ CompactSpace X) where eq_targetAffineLocally' := by ext X Y f diff --git a/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean b/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean index a7b0a7bbd4948..19388a0c54fbd 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean @@ -107,7 +107,6 @@ theorem quasiCompact_affineProperty_iff_quasiSeparatedSpace {X Y : Scheme} [IsAf theorem quasiSeparated_eq_diagonal_is_quasiCompact : @QuasiSeparated = MorphismProperty.diagonal @QuasiCompact := by ext; exact quasiSeparated_iff _ -@[reducible] instance : HasAffineProperty @QuasiSeparated (fun X _ _ _ ↦ QuasiSeparatedSpace X) where __ := HasAffineProperty.copy quasiSeparated_eq_diagonal_is_quasiCompact.symm diff --git a/Mathlib/AlgebraicGeometry/Spec.lean b/Mathlib/AlgebraicGeometry/Spec.lean index b39f3a28bab75..b16f662644390 100644 --- a/Mathlib/AlgebraicGeometry/Spec.lean +++ b/Mathlib/AlgebraicGeometry/Spec.lean @@ -241,10 +241,10 @@ def Spec.locallyRingedSpaceMap {R S : CommRingCat.{u}} (f : R ⟶ S) : replace ha := (isUnit_map_iff (stalkIso S p).inv _).mp ha -- Porting note: `f` had to be made explicit replace ha := IsLocalHom.map_nonunit - (f := (Localization.localRingHom (PrimeSpectrum.comap f p).asIdeal p.asIdeal f _)) _ ha + (f := (Localization.localRingHom (PrimeSpectrum.comap f p).asIdeal p.asIdeal f _)) + ((stalkIso R ((PrimeSpectrum.comap f) p)).hom a) ha convert RingHom.isUnit_map (stalkIso R (PrimeSpectrum.comap f p)).inv ha - erw [← comp_apply, show stalkToFiberRingHom R _ = (stalkIso _ _).hom from rfl, - Iso.hom_inv_id, id_apply] + erw [← comp_apply, Iso.hom_inv_id, id_apply] @[simp] theorem Spec.locallyRingedSpaceMap_id (R : CommRingCat.{u}) : diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean index 6cc0fe177c661..103afe35d71f9 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean @@ -248,7 +248,7 @@ lemma mem_Icc_algebraMap_iff_nnnorm_le {x : A} {r : ℝ≥0} : lemma mem_Icc_iff_norm_le_one {x : A} : x ∈ Icc 0 1 ↔ 0 ≤ x ∧ ‖x‖ ≤ 1 := by - simpa only [map_one] using mem_Icc_algebraMap_iff_norm_le zero_le_one + simpa only [map_one] using mem_Icc_algebraMap_iff_norm_le zero_le_one (A := A) lemma mem_Icc_iff_nnnorm_le_one {x : A} : x ∈ Icc 0 1 ↔ 0 ≤ x ∧ ‖x‖₊ ≤ 1 := @@ -291,9 +291,10 @@ lemma le_iff_norm_sqrt_mul_rpow {a b : A} (hbu : IsUnit b) (ha : 0 ≤ a) (hb : lift b to Aˣ using hbu have hbab : 0 ≤ (b : A) ^ (-(1 / 2) : ℝ) * a * (b : A) ^ (-(1 / 2) : ℝ) := conjugate_nonneg_of_nonneg ha rpow_nonneg + #adaptation_note /-- 2024-11-10 added `(R := A)` -/ conv_rhs => rw [← sq_le_one_iff₀ (norm_nonneg _), sq, ← CStarRing.norm_star_mul_self, star_mul, - IsSelfAdjoint.of_nonneg sqrt_nonneg, IsSelfAdjoint.of_nonneg rpow_nonneg, + IsSelfAdjoint.of_nonneg (R := A) sqrt_nonneg, IsSelfAdjoint.of_nonneg rpow_nonneg, ← mul_assoc, mul_assoc _ _ (sqrt a), sqrt_mul_sqrt_self a, CStarAlgebra.norm_le_one_iff_of_nonneg _ hbab] refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ diff --git a/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean b/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean index a022b455054dd..06469937e9249 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean @@ -207,7 +207,7 @@ noncomputable def equivSigma : ((i : Fin c.length) × Fin (c.partSize i)) ≃ Fi lemma length_pos (h : 0 < n) : 0 < c.length := Nat.zero_lt_of_lt (c.index ⟨0, h⟩).2 lemma neZero_length [NeZero n] (c : OrderedFinpartition n) : NeZero c.length := - ⟨(c.length_pos size_pos').ne'⟩ + ⟨(c.length_pos pos').ne'⟩ lemma neZero_partSize (c : OrderedFinpartition n) (i : Fin c.length) : NeZero (c.partSize i) := .of_pos (c.partSize_pos i) diff --git a/Mathlib/Analysis/Calculus/FDeriv/Mul.lean b/Mathlib/Analysis/Calculus/FDeriv/Mul.lean index e8073446374b3..3bc95aa8d8223 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Mul.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Mul.lean @@ -678,7 +678,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] (List.getElem_mem i.isLt))) ?_ + (hasFDerivAt_pi.mpr fun i ↦ h l[i] (l.getElem_mem _))) ?_ ext m simp [← List.map_map] @@ -690,7 +690,7 @@ theorem HasFDerivWithinAt.list_prod' {l : List ι} {x : E} smulRight (f' l[i]) ((l.drop (.succ i)).map (f · x)).prod) s x := by simp only [← List.finRange_map_get l, List.map_map] refine .congr_fderiv (hasFDerivAt_list_prod_finRange'.comp_hasFDerivWithinAt x - (hasFDerivWithinAt_pi.mpr fun i ↦ h l[i] (l.get_mem i i.isLt))) ?_ + (hasFDerivWithinAt_pi.mpr fun i ↦ h l[i] (l.getElem_mem _))) ?_ ext m simp [← List.map_map] diff --git a/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean b/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean index f57cd0a063cfd..ad88a9cc37573 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean @@ -417,6 +417,8 @@ theorem second_derivative_symmetric_of_eventually {f' : E → E →L[𝕜] F} {x let _ := IsRCLikeNormedField.rclike 𝕜 let _ : NormedSpace ℝ E := NormedSpace.restrictScalars ℝ 𝕜 E let _ : NormedSpace ℝ F := NormedSpace.restrictScalars ℝ 𝕜 F + let _ : LinearMap.CompatibleSMul E F ℝ 𝕜 := LinearMap.IsScalarTower.compatibleSMul + let _ : LinearMap.CompatibleSMul E (E →L[𝕜] F) ℝ 𝕜 := LinearMap.IsScalarTower.compatibleSMul let f'R : E → E →L[ℝ] F := fun x ↦ (f' x).restrictScalars ℝ have hfR : ∀ᶠ y in 𝓝 x, HasFDerivAt f (f'R y) y := by filter_upwards [hf] with y hy using HasFDerivAt.restrictScalars ℝ hy diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean index 88a6d6642a9b9..c92c0c0ce665b 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean @@ -157,8 +157,14 @@ theorem cmp_dist_eq_cmp_dist_coe_center (z w : ℍ) (r : ℝ) : ((mul_neg_of_pos_of_neg w.im_pos (sinh_neg_iff.2 hr₀)).trans_le dist_nonneg).cmp_eq_gt.symm] have hr₀' : 0 ≤ w.im * Real.sinh r := by positivity have hzw₀ : 0 < 2 * z.im * w.im := by positivity - simp only [← cosh_strictMonoOn.cmp_map_eq dist_nonneg hr₀, ← - (pow_left_strictMonoOn₀ two_ne_zero).cmp_map_eq dist_nonneg hr₀', dist_coe_center_sq] + #adaptation_note + /-- + After the bug fix in https://github.com/leanprover/lean4/pull/6024, + we need to give Lean the hint `(y := w.im * Real.sinh r)`. + -/ + simp only [← cosh_strictMonoOn.cmp_map_eq dist_nonneg hr₀, + ← (pow_left_strictMonoOn₀ two_ne_zero).cmp_map_eq dist_nonneg (y := w.im * Real.sinh r) hr₀', + dist_coe_center_sq] rw [← cmp_mul_pos_left hzw₀, ← cmp_sub_zero, ← mul_sub, ← cmp_add_right, zero_add] theorem dist_eq_iff_dist_coe_center_eq : diff --git a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean index 12f3131cc2e8b..e18e2e98d3690 100644 --- a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean +++ b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean @@ -680,7 +680,10 @@ theorem eLpNorm_le_eLpNorm_fderiv_of_le [FiniteDimensional ℝ F] have hp' : p'⁻¹ = p⁻¹ - (finrank ℝ E : ℝ)⁻¹ := by rw [inv_inv, NNReal.coe_sub] · simp - · gcongr + · #adaptation_note + /-- This should just be `gcongr`, but this is not working as of nightly-2024-11-20. + Possibly related to #19262 (since this proof fails at `with_reducible_and_instances`). -/ + exact inv_anti₀ (by positivity) h2p.le have : (q : ℝ≥0∞) ≤ p' := by have H : (p' : ℝ)⁻¹ ≤ (↑q)⁻¹ := trans hp' hpq norm_cast at H ⊢ @@ -688,7 +691,10 @@ theorem eLpNorm_le_eLpNorm_fderiv_of_le [FiniteDimensional ℝ F] · dsimp have : 0 < p⁻¹ - (finrank ℝ E : ℝ≥0)⁻¹ := by simp only [tsub_pos_iff_lt] - gcongr + #adaptation_note + /-- This should just be `gcongr`, but this is not working as of nightly-2024-11-20. + Possibly related to #19262 (since this proof fails at `with_reducible_and_instances`). -/ + exact inv_strictAnti₀ (by positivity) h2p positivity · positivity set t := (μ s).toNNReal ^ (1 / q - 1 / p' : ℝ) diff --git a/Mathlib/Analysis/Normed/Operator/BoundedLinearMaps.lean b/Mathlib/Analysis/Normed/Operator/BoundedLinearMaps.lean index 33d288c85a3c2..9ad44b68f8091 100644 --- a/Mathlib/Analysis/Normed/Operator/BoundedLinearMaps.lean +++ b/Mathlib/Analysis/Normed/Operator/BoundedLinearMaps.lean @@ -190,15 +190,22 @@ theorem isBoundedLinearMap_prod_multilinear {E : ι → Type*} [∀ i, Seminorme [∀ i, NormedSpace 𝕜 (E i)] : IsBoundedLinearMap 𝕜 fun p : ContinuousMultilinearMap 𝕜 E F × ContinuousMultilinearMap 𝕜 E G => p.1.prod p.2 := - (ContinuousMultilinearMap.prodL ..).toContinuousLinearEquiv + (ContinuousMultilinearMap.prodL 𝕜 E F G).toContinuousLinearEquiv |>.toContinuousLinearMap.isBoundedLinearMap +#adaptation_note +/-- +After https://github.com/leanprover/lean4/pull/6024 +we needed to add the named arguments `(ι := ι) (G := F)` +to `ContinuousMultilinearMap.compContinuousLinearMapL`. +-/ /-- Given a fixed continuous linear map `g`, associating to a continuous multilinear map `f` the continuous multilinear map `f (g m₁, ..., g mₙ)` is a bounded linear operation. -/ theorem isBoundedLinearMap_continuousMultilinearMap_comp_linear (g : G →L[𝕜] E) : IsBoundedLinearMap 𝕜 fun f : ContinuousMultilinearMap 𝕜 (fun _ : ι => E) F => f.compContinuousLinearMap fun _ => g := - (ContinuousMultilinearMap.compContinuousLinearMapL fun _ ↦ g).isBoundedLinearMap + (ContinuousMultilinearMap.compContinuousLinearMapL (ι := ι) (G := F) (fun _ ↦ g)) + |>.isBoundedLinearMap end diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean index 2e44424c0033c..ca6f026c41a7c 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean @@ -541,7 +541,7 @@ search. -/ instance normedSpace' : NormedSpace 𝕜' (ContinuousMultilinearMap 𝕜 (fun _ : ι => G') G) := ContinuousMultilinearMap.normedSpace -@[deprecated (since := "2024-11-24")] +@[deprecated norm_neg (since := "2024-11-24")] theorem opNorm_neg (f : ContinuousMultilinearMap 𝕜 E G) : ‖-f‖ = ‖f‖ := norm_neg f @[deprecated (since := "2024-02-02")] alias op_norm_neg := norm_neg diff --git a/Mathlib/Analysis/SumOverResidueClass.lean b/Mathlib/Analysis/SumOverResidueClass.lean index d0109a33064e6..eb5a5d4f11043 100644 --- a/Mathlib/Analysis/SumOverResidueClass.lean +++ b/Mathlib/Analysis/SumOverResidueClass.lean @@ -64,7 +64,7 @@ lemma not_summable_indicator_mod_of_antitone_of_neg {m : ℕ} [hm : NeZero m] {f rw [← ZMod.natCast_zmod_val k, summable_indicator_mod_iff_summable] exact not_summable_of_antitone_of_neg (hf.comp_monotone <| (Covariant.monotone_of_const m).add_const k.val) <| - (hf <| (Nat.le_mul_of_pos_left n Fin.size_pos').trans <| Nat.le_add_right ..).trans_lt hn + (hf <| (Nat.le_mul_of_pos_left n Fin.pos').trans <| Nat.le_add_right ..).trans_lt hn /-- If a decreasing sequence of real numbers is summable on one residue class modulo `m`, then it is also summable on every other residue class mod `m`. -/ diff --git a/Mathlib/CategoryTheory/GlueData.lean b/Mathlib/CategoryTheory/GlueData.lean index 936283e96bec5..5f09d792ec41d 100644 --- a/Mathlib/CategoryTheory/GlueData.lean +++ b/Mathlib/CategoryTheory/GlueData.lean @@ -443,7 +443,7 @@ def GlueData'.t'' (D : GlueData' C) (i j k : D.J) : else haveI := Ne.symm hij pullback.map _ _ _ _ (eqToHom (by aesop)) (eqToHom (by rw [dif_neg hik])) - (eqToHom (by aesop)) (by aesop) (by delta f'; aesop) ≫ + (eqToHom (by aesop)) (by delta f'; aesop) (by delta f'; aesop) ≫ D.t' i j k hij hik hjk ≫ pullback.map _ _ _ _ (eqToHom (by aesop)) (eqToHom (by aesop)) (eqToHom (by aesop)) (by delta f'; aesop) (by delta f'; aesop) diff --git a/Mathlib/CategoryTheory/Monoidal/Mod_.lean b/Mathlib/CategoryTheory/Monoidal/Mod_.lean index ec6327ca22917..da996f019d255 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mod_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mod_.lean @@ -88,6 +88,14 @@ def forget : Mod_ A ⥤ C where open CategoryTheory.MonoidalCategory +#adaptation_note +/-- +After https://github.com/leanprover/lean4/pull/6053 +we needed to increase the `maxHeartbeats` limit. + +This may indicate a configuration problem in Aesop. +-/ +set_option maxHeartbeats 400000 in /-- A morphism of monoid objects induces a "restriction" or "comap" functor between the categories of module objects. -/ diff --git a/Mathlib/CategoryTheory/Monoidal/Mon_.lean b/Mathlib/CategoryTheory/Monoidal/Mon_.lean index 90677e4b70360..5669525f6ab4e 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mon_.lean @@ -246,6 +246,14 @@ namespace CategoryTheory.Functor variable {C} {D : Type u₂} [Category.{v₂} D] [MonoidalCategory.{v₂} D] +#adaptation_note +/-- +After https://github.com/leanprover/lean4/pull/6053 +we needed to increase the `maxHeartbeats` limit. + +This may indicate a configuration problem in Aesop. +-/ +set_option maxHeartbeats 400000 in -- TODO: mapMod F A : Mod A ⥤ Mod (F.mapMon A) /-- A lax monoidal functor takes monoid objects to monoid objects. diff --git a/Mathlib/CategoryTheory/Sites/Adjunction.lean b/Mathlib/CategoryTheory/Sites/Adjunction.lean index 8b4192d40635c..eb19bd7cd03c8 100644 --- a/Mathlib/CategoryTheory/Sites/Adjunction.lean +++ b/Mathlib/CategoryTheory/Sites/Adjunction.lean @@ -100,8 +100,8 @@ variable [HasWeakSheafify J D] [ConcreteCategory D] [HasSheafCompose J (forget D @[deprecated (since := "2024-11-26")] alias composeAndSheafifyFromTypes := composeAndSheafify -/-- The adjunction `composeAndSheafify J G ⊣ sheafForget J`. (Use `Sheaf.adjunction`.)-/ -@[deprecated (since := "2024-11-26")] abbrev adjunctionToTypes +/-- The adjunction `composeAndSheafify J G ⊣ sheafForget J`. -/ +@[deprecated Sheaf.adjunction (since := "2024-11-26")] abbrev adjunctionToTypes {G : Type max v₁ u₁ ⥤ D} (adj : G ⊣ forget D) : composeAndSheafify J G ⊣ sheafForget J := adjunction _ adj diff --git a/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean b/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean index fe073269bb474..99b4e12ea6411 100644 --- a/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean +++ b/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean @@ -177,14 +177,14 @@ theorem ruzsa_triangle_inequality_mul_mul_mul (A B C : Finset G) : rw [mem_erase, mem_powerset, ← nonempty_iff_ne_empty] at hU refine cast_le.1 (?_ : (_ : ℚ≥0) ≤ _) push_cast - refine (le_div_iff₀ <| cast_pos.2 hB.card_pos).1 ?_ - rw [mul_div_right_comm, mul_comm _ B] + rw [← le_div_iff₀ (cast_pos.2 hB.card_pos), mul_div_right_comm, mul_comm _ B] refine (Nat.cast_le.2 <| card_le_card_mul_left hU.1).trans ?_ refine le_trans ?_ (mul_le_mul (hUA _ hB') (cast_le.2 <| card_le_card <| mul_subset_mul_right hU.2) (zero_le _) (zero_le _)) - rw [← mul_div_right_comm, ← mul_assoc] - refine (le_div_iff₀ <| cast_pos.2 hU.1.card_pos).2 ?_ + #adaptation_note /-- 2024-11-01 `le_div_iff₀` is synthesizing wrong `GroupWithZero` without `@` -/ + rw [← mul_div_right_comm, ← mul_assoc, + @le_div_iff₀ _ (_) _ _ _ _ _ _ _ (cast_pos.2 hU.1.card_pos)] exact mod_cast pluennecke_petridis_inequality_mul C (mul_aux hU.1 hU.2 hUA) /-- **Ruzsa's triangle inequality**. Mul-div-div version. -/ diff --git a/Mathlib/Combinatorics/Additive/RuzsaCovering.lean b/Mathlib/Combinatorics/Additive/RuzsaCovering.lean index 825ab42bd37ab..3bbb14e93b8e0 100644 --- a/Mathlib/Combinatorics/Additive/RuzsaCovering.lean +++ b/Mathlib/Combinatorics/Additive/RuzsaCovering.lean @@ -49,8 +49,12 @@ theorem ruzsa_covering_mul (hB : B.Nonempty) (hK : #(A * B) ≤ K * #B) : obtain ⟨b, hb, c, hc₁, hc₂⟩ := H exact mem_mul.2 ⟨b, hb, b⁻¹ * a, mem_div.2 ⟨_, hc₂, _, hc₁, by simp⟩, by simp⟩ -@[to_additive (attr := deprecated (since := "2024-11-26"))] +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +@[to_additive] alias exists_subset_mul_div := ruzsa_covering_mul +attribute [deprecated ruzsa_covering_mul (since := "2024-11-26")] exists_subset_mul_div +attribute [deprecated ruzsa_covering_add (since := "2024-11-26")] exists_subset_add_sub end Finset @@ -68,7 +72,11 @@ lemma ruzsa_covering_mul (hA : A.Finite) (hB : B.Finite) (hB₀ : B.Nonempty) obtain ⟨F, hFA, hF, hAF⟩ := Finset.ruzsa_covering_mul hB₀ (by simpa [← Finset.coe_mul] using hK) exact ⟨F, by norm_cast; simp [*]⟩ -@[to_additive (attr := deprecated (since := "2024-11-26"))] +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +@[to_additive] alias exists_subset_mul_div := ruzsa_covering_mul +attribute [deprecated ruzsa_covering_mul (since := "2024-11-26")] exists_subset_mul_div +attribute [deprecated ruzsa_covering_add (since := "2024-11-26")] exists_subset_add_sub end Set diff --git a/Mathlib/Combinatorics/Colex.lean b/Mathlib/Combinatorics/Colex.lean index b828869fcbf1e..7c732db7e7473 100644 --- a/Mathlib/Combinatorics/Colex.lean +++ b/Mathlib/Combinatorics/Colex.lean @@ -534,7 +534,7 @@ theorem geomSum_injective {n : ℕ} (hn : 2 ≤ n) : geomSum_le_geomSum_iff_toColex_le_toColex hn, ← le_antisymm_iff, Colex.toColex.injEq] at h theorem lt_geomSum_of_mem {a : ℕ} (hn : 2 ≤ n) (hi : a ∈ s) : a < ∑ i in s, n ^ i := - (Nat.lt_pow_self hn a).trans_le <| single_le_sum (by simp) hi + (a.lt_pow_self hn).trans_le <| single_le_sum (by simp) hi @[simp] theorem toFinset_bitIndices_twoPowSum (s : Finset ℕ) : (∑ i in s, 2 ^ i).bitIndices.toFinset = s := by diff --git a/Mathlib/Combinatorics/Enumerative/Composition.lean b/Mathlib/Combinatorics/Enumerative/Composition.lean index 986e40eed5b1b..e1bf3970d658a 100644 --- a/Mathlib/Combinatorics/Enumerative/Composition.lean +++ b/Mathlib/Combinatorics/Enumerative/Composition.lean @@ -150,7 +150,7 @@ theorem sum_blocksFun : ∑ i, c.blocksFun i = n := by conv_rhs => rw [← c.blocks_sum, ← ofFn_blocksFun, sum_ofFn] theorem blocksFun_mem_blocks (i : Fin c.length) : c.blocksFun i ∈ c.blocks := - get_mem _ _ _ + get_mem _ _ @[simp] theorem one_le_blocks {i : ℕ} (h : i ∈ c.blocks) : 1 ≤ i := @@ -158,7 +158,7 @@ theorem one_le_blocks {i : ℕ} (h : i ∈ c.blocks) : 1 ≤ i := @[simp] theorem one_le_blocks' {i : ℕ} (h : i < c.length) : 1 ≤ c.blocks[i] := - c.one_le_blocks (get_mem (blocks c) i h) + c.one_le_blocks (get_mem (blocks c) _) @[simp] theorem blocks_pos' (i : ℕ) (h : i < c.length) : 0 < c.blocks[i] := diff --git a/Mathlib/Computability/Ackermann.lean b/Mathlib/Computability/Ackermann.lean index ace0d47e3ed95..92b4e6f1c3a1a 100644 --- a/Mathlib/Computability/Ackermann.lean +++ b/Mathlib/Computability/Ackermann.lean @@ -235,7 +235,7 @@ private theorem sq_le_two_pow_add_one_minus_three (n : ℕ) : n ^ 2 ≤ 2 ^ (n + norm_num apply succ_le_of_lt rw [Nat.pow_succ, mul_comm _ 2, mul_lt_mul_left (zero_lt_two' ℕ)] - apply lt_two_pow + exact Nat.lt_two_pow_self · rw [Nat.pow_succ, Nat.pow_succ] linarith [one_le_pow k 2 zero_lt_two] diff --git a/Mathlib/Computability/Halting.lean b/Mathlib/Computability/Halting.lean index 0bce012c8ecc0..7437dbc1b31ac 100644 --- a/Mathlib/Computability/Halting.lean +++ b/Mathlib/Computability/Halting.lean @@ -267,11 +267,12 @@ namespace Nat open Vector Part /-- A simplified basis for `Partrec`. -/ -inductive Partrec' : ∀ {n}, (Vector ℕ n →. ℕ) → Prop +inductive Partrec' : ∀ {n}, (Mathlib.Vector ℕ n →. ℕ) → Prop | prim {n f} : @Primrec' n f → @Partrec' n f - | comp {m n f} (g : Fin n → Vector ℕ m →. ℕ) : - Partrec' f → (∀ i, Partrec' (g i)) → Partrec' fun v => (Vector.mOfFn fun i => g i v) >>= f - | rfind {n} {f : Vector ℕ (n + 1) → ℕ} : + | comp {m n f} (g : Fin n → Mathlib.Vector ℕ m →. ℕ) : + Partrec' f → (∀ i, Partrec' (g i)) → + Partrec' fun v => (Mathlib.Vector.mOfFn fun i => g i v) >>= f + | rfind {n} {f : Mathlib.Vector ℕ (n + 1) → ℕ} : @Partrec' (n + 1) f → Partrec' fun v => rfind fun n => some (f (n ::ᵥ v) = 0) end Nat @@ -293,10 +294,11 @@ theorem to_part {n f} (pf : @Partrec' n f) : _root_.Partrec f := by this).to₂.partrec₂ exact _root_.Partrec.rfind this -theorem of_eq {n} {f g : Vector ℕ n →. ℕ} (hf : Partrec' f) (H : ∀ i, f i = g i) : Partrec' g := +theorem of_eq {n} {f g : Mathlib.Vector ℕ n →. ℕ} (hf : Partrec' f) (H : ∀ i, f i = g i) : + Partrec' g := (funext H : f = g) ▸ hf -theorem of_prim {n} {f : Vector ℕ n → ℕ} (hf : Primrec f) : @Partrec' n f := +theorem of_prim {n} {f : Mathlib.Vector ℕ n → ℕ} (hf : Primrec f) : @Partrec' n f := prim (Nat.Primrec'.of_prim hf) theorem head {n : ℕ} : @Partrec' n.succ (@head ℕ n) := @@ -313,21 +315,21 @@ protected theorem bind {n f g} (hf : @Partrec' n f) (hg : @Partrec' (n + 1) g) : exact prim (Nat.Primrec'.get _)).of_eq fun v => by simp [mOfFn, Part.bind_assoc, pure] -protected theorem map {n f} {g : Vector ℕ (n + 1) → ℕ} (hf : @Partrec' n f) +protected theorem map {n f} {g : Mathlib.Vector ℕ (n + 1) → ℕ} (hf : @Partrec' n f) (hg : @Partrec' (n + 1) g) : @Partrec' n fun v => (f v).map fun a => g (a ::ᵥ v) := by simpa [(Part.bind_some_eq_map _ _).symm] using hf.bind hg /-- Analogous to `Nat.Partrec'` for `ℕ`-valued functions, a predicate for partial recursive vector-valued functions. -/ -def Vec {n m} (f : Vector ℕ n → Vector ℕ m) := +def Vec {n m} (f : Mathlib.Vector ℕ n → Mathlib.Vector ℕ m) := ∀ i, Partrec' fun v => (f v).get i nonrec theorem Vec.prim {n m f} (hf : @Nat.Primrec'.Vec n m f) : Vec f := fun i => prim <| hf i protected theorem nil {n} : @Vec n 0 fun _ => nil := fun i => i.elim0 -protected theorem cons {n m} {f : Vector ℕ n → ℕ} {g} (hf : @Partrec' n f) (hg : @Vec n m g) : - Vec fun v => f v ::ᵥ g v := fun i => +protected theorem cons {n m} {f : Mathlib.Vector ℕ n → ℕ} {g} (hf : @Partrec' n f) + (hg : @Vec n m g) : Vec fun v => f v ::ᵥ g v := fun i => Fin.cases (by simpa using hf) (fun i => by simp only [hg i, get_cons_succ]) i theorem idv {n} : @Vec n n id := @@ -336,11 +338,11 @@ theorem idv {n} : @Vec n n id := theorem comp' {n m f g} (hf : @Partrec' m f) (hg : @Vec n m g) : Partrec' fun v => f (g v) := (hf.comp _ hg).of_eq fun v => by simp -theorem comp₁ {n} (f : ℕ →. ℕ) {g : Vector ℕ n → ℕ} (hf : @Partrec' 1 fun v => f v.head) +theorem comp₁ {n} (f : ℕ →. ℕ) {g : Mathlib.Vector ℕ n → ℕ} (hf : @Partrec' 1 fun v => f v.head) (hg : @Partrec' n g) : @Partrec' n fun v => f (g v) := by simpa using hf.comp' (Partrec'.cons hg Partrec'.nil) -theorem rfindOpt {n} {f : Vector ℕ (n + 1) → ℕ} (hf : @Partrec' (n + 1) f) : +theorem rfindOpt {n} {f : Mathlib.Vector ℕ (n + 1) → ℕ} (hf : @Partrec' (n + 1) f) : @Partrec' n fun v => Nat.rfindOpt fun a => ofNat (Option ℕ) (f (a ::ᵥ v)) := ((rfind <| (of_prim (Primrec.nat_sub.comp (_root_.Primrec.const 1) Primrec.vector_head)).comp₁ @@ -367,7 +369,7 @@ open Nat.Partrec.Code theorem of_part : ∀ {n f}, _root_.Partrec f → @Partrec' n f := @(suffices ∀ f, Nat.Partrec f → @Partrec' 1 fun v => f v.head from fun {n f} hf => by let g := fun n₁ => - (Part.ofOption (decode (α := Vector ℕ n) n₁)).bind (fun a => Part.map encode (f a)) + (Part.ofOption (decode (α := Mathlib.Vector ℕ n) n₁)).bind (fun a => Part.map encode (f a)) exact (comp₁ g (this g hf) (prim Nat.Primrec'.encode)).of_eq fun i => by dsimp only [g]; simp [encodek, Part.map_id'] diff --git a/Mathlib/Computability/Partrec.lean b/Mathlib/Computability/Partrec.lean index 700990868e922..3ad33d20e77c5 100644 --- a/Mathlib/Computability/Partrec.lean +++ b/Mathlib/Computability/Partrec.lean @@ -200,7 +200,7 @@ theorem ppred : Partrec fun n => ppred n := eq_none_iff.2 fun a ⟨⟨m, h, _⟩, _⟩ => by simp [show 0 ≠ m.succ by intro h; injection h] at h · refine eq_some_iff.2 ?_ - simp only [mem_rfind, not_true, IsEmpty.forall_iff, decide_True, mem_some_iff, + simp only [mem_rfind, not_true, IsEmpty.forall_iff, decide_true, mem_some_iff, false_eq_decide_iff, true_and] intro m h simp [ne_of_gt h] @@ -314,25 +314,25 @@ theorem list_concat : Computable₂ fun l (a : α) => l ++ [a] := theorem list_length : Computable (@List.length α) := Primrec.list_length.to_comp -theorem vector_cons {n} : Computable₂ (@Vector.cons α n) := +theorem vector_cons {n} : Computable₂ (@Mathlib.Vector.cons α n) := Primrec.vector_cons.to_comp -theorem vector_toList {n} : Computable (@Vector.toList α n) := +theorem vector_toList {n} : Computable (@Mathlib.Vector.toList α n) := Primrec.vector_toList.to_comp -theorem vector_length {n} : Computable (@Vector.length α n) := +theorem vector_length {n} : Computable (@Mathlib.Vector.length α n) := Primrec.vector_length.to_comp -theorem vector_head {n} : Computable (@Vector.head α n) := +theorem vector_head {n} : Computable (@Mathlib.Vector.head α n) := Primrec.vector_head.to_comp -theorem vector_tail {n} : Computable (@Vector.tail α n) := +theorem vector_tail {n} : Computable (@Mathlib.Vector.tail α n) := Primrec.vector_tail.to_comp -theorem vector_get {n} : Computable₂ (@Vector.get α n) := +theorem vector_get {n} : Computable₂ (@Mathlib.Vector.get α n) := Primrec.vector_get.to_comp -theorem vector_ofFn' {n} : Computable (@Vector.ofFn α n) := +theorem vector_ofFn' {n} : Computable (@Mathlib.Vector.ofFn α n) := Primrec.vector_ofFn'.to_comp theorem fin_app {n} : Computable₂ (@id (Fin n → σ)) := @@ -523,7 +523,8 @@ end Partrec @[simp] theorem Vector.mOfFn_part_some {α n} : - ∀ f : Fin n → α, (Vector.mOfFn fun i => Part.some (f i)) = Part.some (Vector.ofFn f) := + ∀ f : Fin n → α, + (Mathlib.Vector.mOfFn fun i => Part.some (f i)) = Part.some (Mathlib.Vector.ofFn f) := Vector.mOfFn_pure namespace Computable @@ -642,7 +643,7 @@ theorem list_ofFn : exact list_cons.comp (hf 0) (list_ofFn fun i => hf i.succ) theorem vector_ofFn {n} {f : Fin n → α → σ} (hf : ∀ i, Computable (f i)) : - Computable fun a => Vector.ofFn fun i => f i a := + Computable fun a => Mathlib.Vector.ofFn fun i => f i a := (Partrec.vector_mOfFn hf).of_eq fun a => by simp end Computable diff --git a/Mathlib/Computability/PartrecCode.lean b/Mathlib/Computability/PartrecCode.lean index df168bf2c9f6d..02daf224ca8e5 100644 --- a/Mathlib/Computability/PartrecCode.lean +++ b/Mathlib/Computability/PartrecCode.lean @@ -902,7 +902,7 @@ private theorem evaln_map (k c n) : ((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] + · rw [List.getElem?_eq_none] · cases e : evaln k c n · rfl exact kn.elim (evaln_bound e) diff --git a/Mathlib/Computability/Primrec.lean b/Mathlib/Computability/Primrec.lean index f7492a1e781aa..76aceca0eb52f 100644 --- a/Mathlib/Computability/Primrec.lean +++ b/Mathlib/Computability/Primrec.lean @@ -1101,7 +1101,7 @@ def subtype {p : α → Prop} [DecidablePred p] (hp : PrimrecPred p) : Primcodab instance fin {n} : Primcodable (Fin n) := @ofEquiv _ _ (subtype <| nat_lt.comp .id (const n)) Fin.equivSubtype -instance vector {n} : Primcodable (Vector α n) := +instance vector {n} : Primcodable (Mathlib.Vector α n) := subtype ((@Primrec.eq ℕ _ _).comp list_length (const _)) instance finArrow {n} : Primcodable (Fin n → α) := @@ -1184,25 +1184,26 @@ theorem fin_val {n} : Primrec (fun (i : Fin n) => (i : ℕ)) := theorem fin_succ {n} : Primrec (@Fin.succ n) := fin_val_iff.1 <| by simp [succ.comp fin_val] -theorem vector_toList {n} : Primrec (@Vector.toList α n) := +theorem vector_toList {n} : Primrec (@Mathlib.Vector.toList α n) := subtype_val -theorem vector_toList_iff {n} {f : α → Vector β n} : (Primrec fun a => (f a).toList) ↔ Primrec f := +theorem vector_toList_iff {n} {f : α → Mathlib.Vector β n} : + (Primrec fun a => (f a).toList) ↔ Primrec f := subtype_val_iff -theorem vector_cons {n} : Primrec₂ (@Vector.cons α n) := +theorem vector_cons {n} : Primrec₂ (@Mathlib.Vector.cons α n) := vector_toList_iff.1 <| by simpa using list_cons.comp fst (vector_toList_iff.2 snd) -theorem vector_length {n} : Primrec (@Vector.length α n) := +theorem vector_length {n} : Primrec (@Mathlib.Vector.length α n) := const _ -theorem vector_head {n} : Primrec (@Vector.head α n) := +theorem vector_head {n} : Primrec (@Mathlib.Vector.head α n) := option_some_iff.1 <| (list_head?.comp vector_toList).of_eq fun ⟨_ :: _, _⟩ => rfl -theorem vector_tail {n} : Primrec (@Vector.tail α n) := +theorem vector_tail {n} : Primrec (@Mathlib.Vector.tail α n) := vector_toList_iff.1 <| (list_tail.comp vector_toList).of_eq fun ⟨l, h⟩ => by cases l <;> rfl -theorem vector_get {n} : Primrec₂ (@Vector.get α n) := +theorem vector_get {n} : Primrec₂ (@Mathlib.Vector.get α n) := option_some_iff.1 <| (list_get?.comp (vector_toList.comp fst) (fin_val.comp snd)).of_eq fun a => by rw [Vector.get_eq_get, ← List.get?_eq_get] @@ -1215,13 +1216,13 @@ theorem list_ofFn : simpa [List.ofFn_succ] using list_cons.comp (hf 0) (list_ofFn fun i => hf i.succ) theorem vector_ofFn {n} {f : Fin n → α → σ} (hf : ∀ i, Primrec (f i)) : - Primrec fun a => Vector.ofFn fun i => f i a := + Primrec fun a => Mathlib.Vector.ofFn fun i => f i a := vector_toList_iff.1 <| by simp [list_ofFn hf] -theorem vector_get' {n} : Primrec (@Vector.get α n) := +theorem vector_get' {n} : Primrec (@Mathlib.Vector.get α n) := of_equiv_symm -theorem vector_ofFn' {n} : Primrec (@Vector.ofFn α n) := +theorem vector_ofFn' {n} : Primrec (@Mathlib.Vector.ofFn α n) := of_equiv theorem fin_app {n} : Primrec₂ (@id (Fin n → σ)) := @@ -1248,16 +1249,16 @@ open Mathlib.Vector work with n-ary functions on ℕ instead of unary functions. We prove that this is equivalent to the regular notion in `to_prim` and `of_prim`. -/ -inductive Primrec' : ∀ {n}, (Vector ℕ n → ℕ) → Prop +inductive Primrec' : ∀ {n}, (Mathlib.Vector ℕ n → ℕ) → Prop | zero : @Primrec' 0 fun _ => 0 | succ : @Primrec' 1 fun v => succ v.head | get {n} (i : Fin n) : Primrec' fun v => v.get i - | comp {m n f} (g : Fin n → Vector ℕ m → ℕ) : - Primrec' f → (∀ i, Primrec' (g i)) → Primrec' fun a => f (ofFn fun i => g i a) + | comp {m n f} (g : Fin n → Mathlib.Vector ℕ m → ℕ) : + Primrec' f → (∀ i, Primrec' (g i)) → Primrec' fun a => f (Mathlib.Vector.ofFn fun i => g i a) | prec {n f g} : @Primrec' n f → @Primrec' (n + 2) g → - Primrec' fun v : Vector ℕ (n + 1) => + Primrec' fun v : Mathlib.Vector ℕ (n + 1) => v.head.rec (f v.tail) fun y IH => g (y ::ᵥ IH ::ᵥ v.tail) end Nat @@ -1280,7 +1281,8 @@ theorem to_prim {n f} (pf : @Nat.Primrec' n f) : Primrec f := by Primrec.vector_cons.comp (Primrec.snd.comp .snd) <| (@Primrec.vector_tail _ _ (n + 1)).comp .fst).to₂ -theorem of_eq {n} {f g : Vector ℕ n → ℕ} (hf : Primrec' f) (H : ∀ i, f i = g i) : Primrec' g := +theorem of_eq {n} {f g : Mathlib.Vector ℕ n → ℕ} (hf : Primrec' f) (H : ∀ i, f i = g i) : + Primrec' g := (funext H : f = g) ▸ hf theorem const {n} : ∀ m, @Primrec' n fun _ => m @@ -1295,7 +1297,7 @@ theorem tail {n f} (hf : @Primrec' n f) : @Primrec' n.succ fun v => f v.tail := rw [← ofFn_get v.tail]; congr; funext i; simp /-- A function from vectors to vectors is primitive recursive when all of its projections are. -/ -def Vec {n m} (f : Vector ℕ n → Vector ℕ m) : Prop := +def Vec {n m} (f : Mathlib.Vector ℕ n → Mathlib.Vector ℕ m) : Prop := ∀ i, Primrec' fun v => (f v).get i protected theorem nil {n} : @Vec n 0 fun _ => nil := fun i => i.elim0 @@ -1390,7 +1392,7 @@ theorem unpair₂ {n f} (hf : @Primrec' n f) : @Primrec' n fun v => (f v).unpair theorem of_prim {n f} : Primrec f → @Primrec' n f := suffices ∀ f, Nat.Primrec f → @Primrec' 1 fun v => f v.head from fun hf => (pred.comp₁ _ <| - (this _ hf).comp₁ (fun m => Encodable.encode <| (@decode (Vector ℕ n) _ m).map f) + (this _ hf).comp₁ (fun m => Encodable.encode <| (@decode (Mathlib.Vector ℕ n) _ m).map f) Primrec'.encode).of_eq fun i => by simp [encodek] fun f hf => by diff --git a/Mathlib/Computability/TMToPartrec.lean b/Mathlib/Computability/TMToPartrec.lean index 5ef3c02a32d1c..effb93335081f 100644 --- a/Mathlib/Computability/TMToPartrec.lean +++ b/Mathlib/Computability/TMToPartrec.lean @@ -239,12 +239,14 @@ def prec (f g : Code) : Code := attribute [-simp] Part.bind_eq_bind Part.map_eq_map Part.pure_eq_some -theorem exists_code.comp {m n} {f : Vector ℕ n →. ℕ} {g : Fin n → Vector ℕ m →. ℕ} - (hf : ∃ c : Code, ∀ v : Vector ℕ n, c.eval v.1 = pure <$> f v) - (hg : ∀ i, ∃ c : Code, ∀ v : Vector ℕ m, c.eval v.1 = pure <$> g i v) : - ∃ c : Code, ∀ v : Vector ℕ m, c.eval v.1 = pure <$> ((Vector.mOfFn fun i => g i v) >>= f) := by +theorem exists_code.comp {m n} {f : Mathlib.Vector ℕ n →. ℕ} {g : Fin n → Mathlib.Vector ℕ m →. ℕ} + (hf : ∃ c : Code, ∀ v : Mathlib.Vector ℕ n, c.eval v.1 = pure <$> f v) + (hg : ∀ i, ∃ c : Code, ∀ v : Mathlib.Vector ℕ m, c.eval v.1 = pure <$> g i v) : + ∃ c : Code, ∀ v : Mathlib.Vector ℕ m, + c.eval v.1 = pure <$> ((Mathlib.Vector.mOfFn fun i => g i v) >>= f) := by rsuffices ⟨cg, hg⟩ : - ∃ c : Code, ∀ v : Vector ℕ m, c.eval v.1 = Subtype.val <$> Vector.mOfFn fun i => g i v + ∃ c : Code, ∀ v : Mathlib.Vector ℕ m, + c.eval v.1 = Subtype.val <$> Mathlib.Vector.mOfFn fun i => g i v · obtain ⟨cf, hf⟩ := hf exact ⟨cf.comp cg, fun v => by @@ -259,8 +261,8 @@ theorem exists_code.comp {m n} {f : Vector ℕ n →. ℕ} {g : Fin n → Vector simp [Vector.mOfFn, hg₁, map_bind, seq_bind_eq, bind_assoc, (· ∘ ·), hl] rfl⟩ -theorem exists_code {n} {f : Vector ℕ n →. ℕ} (hf : Nat.Partrec' f) : - ∃ c : Code, ∀ v : Vector ℕ n, c.eval v.1 = pure <$> f v := by +theorem exists_code {n} {f : Mathlib.Vector ℕ n →. ℕ} (hf : Nat.Partrec' f) : + ∃ c : Code, ∀ v : Mathlib.Vector ℕ n, c.eval v.1 = pure <$> f v := by induction hf with | prim hf => induction hf with @@ -1402,8 +1404,8 @@ theorem succ_ok {q s n} {c d : List Γ'} : simp only [TM2.step, trList, trNat.eq_1, Nat.cast_succ, Num.add_one] cases' (n : Num) with a · refine TransGen.head rfl ?_ - simp only [Option.mem_def, TM2.stepAux, elim_main, decide_False, elim_update_main, ne_eq, - Function.update_noteq, elim_rev, elim_update_rev, decide_True, Function.update_same, + simp only [Option.mem_def, TM2.stepAux, elim_main, decide_false, elim_update_main, ne_eq, + Function.update_noteq, elim_rev, elim_update_rev, decide_true, Function.update_same, cond_true, cond_false] convert unrev_ok using 1 simp only [elim_update_rev, elim_rev, elim_main, List.reverseAux_nil, elim_update_main] @@ -1446,7 +1448,7 @@ theorem pred_ok (q₁ q₂ s v) (c d : List Γ') : ∃ s', · simp only [trPosNum, List.singleton_append, List.nil_append] refine TransGen.head rfl ?_ simp only [Option.mem_def, TM2.stepAux, elim_main, List.head?_cons, Option.some.injEq, - decide_False, List.tail_cons, elim_update_main, ne_eq, Function.update_noteq, elim_rev, + decide_false, List.tail_cons, elim_update_main, ne_eq, Function.update_noteq, elim_rev, elim_update_rev, natEnd, Function.update_same, cond_true, cond_false] convert unrev_ok using 2 simp diff --git a/Mathlib/Computability/TuringMachine.lean b/Mathlib/Computability/TuringMachine.lean index 4e3298c29a939..0f1145c0c16ee 100644 --- a/Mathlib/Computability/TuringMachine.lean +++ b/Mathlib/Computability/TuringMachine.lean @@ -1443,7 +1443,7 @@ open TM1 variable {Γ : Type*} theorem exists_enc_dec [Inhabited Γ] [Finite Γ] : - ∃ (n : ℕ) (enc : Γ → Vector Bool n) (dec : Vector Bool n → Γ), + ∃ (n : ℕ) (enc : Γ → Mathlib.Vector Bool n) (dec : Mathlib.Vector Bool n → Γ), enc default = Vector.replicate n false ∧ ∀ a, dec (enc a) = a := by rcases Finite.exists_equiv_fin Γ with ⟨n, ⟨e⟩⟩ letI : DecidableEq Γ := e.decidableEq @@ -1475,13 +1475,13 @@ local notation "Stmt'₁" => Stmt Bool Λ'₁ σ local notation "Cfg'₁" => Cfg Bool Λ'₁ σ /-- Read a vector of length `n` from the tape. -/ -def readAux : ∀ n, (Vector Bool n → Stmt'₁) → Stmt'₁ +def readAux : ∀ n, (Mathlib.Vector Bool n → Stmt'₁) → Stmt'₁ | 0, f => f Vector.nil | i + 1, f => Stmt.branch (fun a _ ↦ a) (Stmt.move Dir.right <| readAux i fun v ↦ f (true ::ᵥ v)) (Stmt.move Dir.right <| readAux i fun v ↦ f (false ::ᵥ v)) -variable {n : ℕ} (enc : Γ → Vector Bool n) (dec : Vector Bool n → Γ) +variable {n : ℕ} (enc : Γ → Mathlib.Vector Bool n) (dec : Mathlib.Vector Bool n → Γ) /-- A move left or right corresponds to `n` moves across the super-cell. -/ def move (d : Dir) (q : Stmt'₁) : Stmt'₁ := @@ -1530,7 +1530,8 @@ theorem supportsStmt_write {S : Finset Λ'₁} {l : List Bool} {q : Stmt'₁} : theorem supportsStmt_read {S : Finset Λ'₁} : ∀ {f : Γ → Stmt'₁}, (∀ a, SupportsStmt S (f a)) → SupportsStmt S (read dec f) := suffices - ∀ (i) (f : Vector Bool i → Stmt'₁), (∀ v, SupportsStmt S (f v)) → SupportsStmt S (readAux i f) + ∀ (i) (f : Mathlib.Vector Bool i → Stmt'₁), + (∀ v, SupportsStmt S (f v)) → SupportsStmt S (readAux i f) from fun hf ↦ this n _ (by intro; simp only [supportsStmt_move, hf]) fun i f hf ↦ by induction' i with i IH; · exact hf _ @@ -2437,7 +2438,7 @@ theorem tr_respects_aux {q v T k} {S : ∀ k, List (Γ k)} obtain ⟨T', hT', hrun⟩ := tr_respects_aux₂ (Λ := Λ) hT o have := hgo.tail' rfl rw [tr, TM1.stepAux, Tape.move_right_n_head, Tape.mk'_nth_nat, addBottom_nth_snd, - stk_nth_val _ (hT k), List.getElem?_len_le (le_of_eq (List.length_reverse _)), + stk_nth_val _ (hT k), List.getElem?_eq_none (le_of_eq (List.length_reverse _)), Option.isNone, cond, hrun, TM1.stepAux] at this obtain ⟨c, gc, rc⟩ := IH hT' refine ⟨c, gc, (this.to₀.trans (tr_respects_aux₃ M _) c (TransGen.head' rfl ?_)).to_reflTransGen⟩ diff --git a/Mathlib/Control/Random.lean b/Mathlib/Control/Random.lean index c83ecbaba193b..fb87c58b85783 100644 --- a/Mathlib/Control/Random.lean +++ b/Mathlib/Control/Random.lean @@ -93,7 +93,7 @@ def randBound (α : Type u) (BoundedRandom.randomR lo hi h : RandGT g _ _) def randFin {n : Nat} [RandomGen g] : RandGT g m (Fin n.succ) := - fun ⟨g⟩ ↦ pure <| randNat g 0 n |>.map Fin.ofNat ULift.up + fun ⟨g⟩ ↦ pure <| randNat g 0 n |>.map (Fin.ofNat' _) ULift.up instance {n : Nat} : Random m (Fin n.succ) where random := randFin diff --git a/Mathlib/Data/Bool/Basic.lean b/Mathlib/Data/Bool/Basic.lean index 1b5e3ff38f835..5f5bfc9ba28c8 100644 --- a/Mathlib/Data/Bool/Basic.lean +++ b/Mathlib/Data/Bool/Basic.lean @@ -103,9 +103,9 @@ theorem coe_xor_iff (a b : Bool) : xor a b ↔ Xor' (a = true) (b = true) := by end -@[deprecated (since := "2024-06-07")] alias decide_True := decide_true_eq_true +@[deprecated (since := "2024-06-07")] alias decide_True := decide_true -@[deprecated (since := "2024-06-07")] alias decide_False := decide_false_eq_false +@[deprecated (since := "2024-06-07")] alias decide_False := decide_false @[deprecated (since := "2024-06-07")] alias coe_decide := decide_eq_true_iff diff --git a/Mathlib/Data/DFinsupp/Sigma.lean b/Mathlib/Data/DFinsupp/Sigma.lean index 4df5f452d103f..10d83e6df6193 100644 --- a/Mathlib/Data/DFinsupp/Sigma.lean +++ b/Mathlib/Data/DFinsupp/Sigma.lean @@ -63,14 +63,20 @@ theorem sigmaCurry_zero [∀ i j, Zero (δ i j)] : @[simp] theorem sigmaCurry_add [∀ i j, AddZeroClass (δ i j)] (f g : Π₀ (i : Σ _, _), δ i.1 i.2) : - sigmaCurry (f + g) = sigmaCurry f + sigmaCurry g := by + #adaptation_note + /-- After https://github.com/leanprover/lean4/pull/6024 + we needed to add the `(_ : Π₀ (i) (j), δ i j)` type annotation. -/ + sigmaCurry (f + g) = (sigmaCurry f + sigmaCurry g : Π₀ (i) (j), δ i j) := by ext (i j) rfl @[simp] theorem sigmaCurry_smul [Monoid γ] [∀ i j, AddMonoid (δ i j)] [∀ i j, DistribMulAction γ (δ i j)] (r : γ) (f : Π₀ (i : Σ _, _), δ i.1 i.2) : - sigmaCurry (r • f) = r • sigmaCurry f := by + #adaptation_note + /-- After https://github.com/leanprover/lean4/pull/6024 + we needed to add the `(_ : Π₀ (i) (j), δ i j)` type annotation. -/ + sigmaCurry (r • f) = (r • sigmaCurry f : Π₀ (i) (j), δ i j) := by ext (i j) rfl diff --git a/Mathlib/Data/ENNReal/Inv.lean b/Mathlib/Data/ENNReal/Inv.lean index bccdde7f1f4ab..0a17941605844 100644 --- a/Mathlib/Data/ENNReal/Inv.lean +++ b/Mathlib/Data/ENNReal/Inv.lean @@ -543,7 +543,7 @@ theorem exists_inv_two_pow_lt (ha : a ≠ 0) : ∃ n : ℕ, 2⁻¹ ^ n < a := by refine ⟨n, lt_trans ?_ hn⟩ rw [← ENNReal.inv_pow, ENNReal.inv_lt_inv] norm_cast - exact n.lt_two_pow + exact n.lt_two_pow_self @[simp, norm_cast] theorem coe_zpow (hr : r ≠ 0) (n : ℤ) : (↑(r ^ n) : ℝ≥0∞) = (r : ℝ≥0∞) ^ n := by diff --git a/Mathlib/Data/Fin/Tuple/Take.lean b/Mathlib/Data/Fin/Tuple/Take.lean index 1ff923b04788e..462351c63a8d1 100644 --- a/Mathlib/Data/Fin/Tuple/Take.lean +++ b/Mathlib/Data/Fin/Tuple/Take.lean @@ -3,7 +3,6 @@ Copyright (c) 2024 Quang Dao. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Quang Dao -/ -import Batteries.Data.List.OfFn import Mathlib.Data.Fin.Tuple.Basic /-! diff --git a/Mathlib/Data/Finset/Defs.lean b/Mathlib/Data/Finset/Defs.lean index 694a1fbcf2f29..daad50dd4cc92 100644 --- a/Mathlib/Data/Finset/Defs.lean +++ b/Mathlib/Data/Finset/Defs.lean @@ -213,6 +213,8 @@ instance partialOrder : PartialOrder (Finset α) where le_trans _ _ _ hst htu _ ha := htu <| hst ha le_antisymm _ _ hst hts := ext fun _ => ⟨@hst _, @hts _⟩ +theorem subset_of_le : s ≤ t → s ⊆ t := id + instance : IsRefl (Finset α) (· ⊆ ·) := show IsRefl (Finset α) (· ≤ ·) by infer_instance diff --git a/Mathlib/Data/Finset/SDiff.lean b/Mathlib/Data/Finset/SDiff.lean index 52b4f4ac894a3..1b6228553e58b 100644 --- a/Mathlib/Data/Finset/SDiff.lean +++ b/Mathlib/Data/Finset/SDiff.lean @@ -111,11 +111,11 @@ theorem sdiff_empty : s \ ∅ = s := @[mono, gcongr] theorem sdiff_subset_sdiff (hst : s ⊆ t) (hvu : v ⊆ u) : s \ u ⊆ t \ v := - sdiff_le_sdiff hst hvu + subset_of_le (sdiff_le_sdiff hst hvu) theorem sdiff_subset_sdiff_iff_subset {r : Finset α} (hs : s ⊆ r) (ht : t ⊆ r) : - r \ s ⊆ r \ t ↔ t ⊆ s := - sdiff_le_sdiff_iff_le hs ht + r \ s ⊆ r \ t ↔ t ⊆ s := by + simpa only [← le_eq_subset] using sdiff_le_sdiff_iff_le hs ht @[simp, norm_cast] theorem coe_sdiff (s₁ s₂ : Finset α) : ↑(s₁ \ s₂) = (s₁ \ s₂ : Set α) := diff --git a/Mathlib/Data/Finset/Sort.lean b/Mathlib/Data/Finset/Sort.lean index 7e7f3fc62d674..d4354b5039a5b 100644 --- a/Mathlib/Data/Finset/Sort.lean +++ b/Mathlib/Data/Finset/Sort.lean @@ -113,7 +113,7 @@ theorem sorted_zero_eq_min'_aux (s : Finset α) (h : 0 < (s.sort (· ≤ ·)).le obtain ⟨i, hi⟩ : ∃ i, l.get i = s.min' H := List.mem_iff_get.1 this rw [← hi] exact (s.sort_sorted (· ≤ ·)).rel_get_of_le (Nat.zero_le i) - · have : l.get ⟨0, h⟩ ∈ s := (Finset.mem_sort (α := α) (· ≤ ·)).1 (List.get_mem l 0 h) + · have : l.get ⟨0, h⟩ ∈ s := (Finset.mem_sort (α := α) (· ≤ ·)).1 (List.get_mem l _) exact s.min'_le _ this theorem sorted_zero_eq_min' {s : Finset α} {h : 0 < (s.sort (· ≤ ·)).length} : @@ -130,7 +130,7 @@ theorem sorted_last_eq_max'_aux (s : Finset α) let l := s.sort (· ≤ ·) apply le_antisymm · have : l.get ⟨(s.sort (· ≤ ·)).length - 1, h⟩ ∈ s := - (Finset.mem_sort (α := α) (· ≤ ·)).1 (List.get_mem l _ h) + (Finset.mem_sort (α := α) (· ≤ ·)).1 (List.get_mem l _) exact s.le_max' _ this · have : s.max' H ∈ l := (Finset.mem_sort (α := α) (· ≤ ·)).mpr (s.max'_mem H) obtain ⟨i, hi⟩ : ∃ i, l.get i = s.max' H := List.mem_iff_get.1 this diff --git a/Mathlib/Data/Fintype/BigOperators.lean b/Mathlib/Data/Fintype/BigOperators.lean index 6c8241d4dd465..0825cf4335415 100644 --- a/Mathlib/Data/Fintype/BigOperators.lean +++ b/Mathlib/Data/Fintype/BigOperators.lean @@ -172,7 +172,8 @@ theorem Fintype.card_fun [DecidableEq α] [Fintype α] [Fintype β] : simp @[simp] -theorem card_vector [Fintype α] (n : ℕ) : Fintype.card (Vector α n) = Fintype.card α ^ n := by +theorem card_vector [Fintype α] (n : ℕ) : + Fintype.card (Mathlib.Vector α n) = Fintype.card α ^ n := by rw [Fintype.ofEquiv_card]; simp /-- It is equivalent to compute the product of a function over `Fin n` or `Finset.range n`. -/ diff --git a/Mathlib/Data/Fintype/Fin.lean b/Mathlib/Data/Fintype/Fin.lean index 8744f61e2f3ae..96b7ea32bd858 100644 --- a/Mathlib/Data/Fintype/Fin.lean +++ b/Mathlib/Data/Fintype/Fin.lean @@ -61,7 +61,7 @@ theorem card_filter_univ_succ' (p : Fin (n + 1) → Prop) [DecidablePred p] : #{x | p x} = ite (p 0) 1 0 + #{x | p (.succ x)}:= by rw [card_filter_univ_succ]; split_ifs <;> simp [add_comm] -theorem card_filter_univ_eq_vector_get_eq_count [DecidableEq α] (a : α) (v : Vector α n) : +theorem card_filter_univ_eq_vector_get_eq_count [DecidableEq α] (a : α) (v : Mathlib.Vector α n) : #{i | v.get i = a} = v.toList.count a := by induction' v with n x xs hxs · simp diff --git a/Mathlib/Data/Fintype/Vector.lean b/Mathlib/Data/Fintype/Vector.lean index 2e94be07c8dbd..9b833c039d387 100644 --- a/Mathlib/Data/Fintype/Vector.lean +++ b/Mathlib/Data/Fintype/Vector.lean @@ -14,7 +14,7 @@ open Mathlib (Vector) variable {α : Type*} -instance Vector.fintype [Fintype α] {n : ℕ} : Fintype (Vector α n) := +instance Vector.fintype [Fintype α] {n : ℕ} : Fintype (Mathlib.Vector α n) := Fintype.ofEquiv _ (Equiv.vectorEquivFin _ _).symm instance [DecidableEq α] [Fintype α] {n : ℕ} : Fintype (Sym.Sym' α n) := by diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 3f51c9f665d12..5b5ca26520e79 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -329,7 +329,7 @@ theorem getLast_replicate_succ (m : ℕ) (a : α) : /-- If the last element of `l` does not satisfy `p`, then it is also the last element of `l.filter p`. -/ -lemma getLast_filter {p : α → Bool} : +lemma getLast_filter' {p : α → Bool} : ∀ (l : List α) (hlp : l.filter p ≠ []), p (l.getLast (hlp <| ·.symm ▸ rfl)) = true → (l.filter p).getLast hlp = l.getLast (hlp <| ·.symm ▸ rfl) | [a], h, h' => by rw [List.getLast_singleton'] at h'; simp [List.filter_cons, h'] @@ -338,10 +338,10 @@ lemma getLast_filter {p : α → Bool} : simp only [List.filter_cons (x := a)] at h ⊢ obtain ha | ha := Bool.eq_false_or_eq_true (p a) · simp only [ha, ite_true] - rw [getLast_cons, getLast_filter (b :: as) _ h'] + rw [getLast_cons, getLast_filter' (b :: as) _ h'] exact ne_nil_of_mem <| mem_filter.2 ⟨getLast_mem _, h'⟩ · simp only [ha, cond_false] at h ⊢ - exact getLast_filter (b :: as) h h' + exact getLast_filter' (b :: as) h h' /-! ### getLast? -/ @@ -739,10 +739,10 @@ end IndexOf section deprecated @[simp] -theorem getElem?_length (l : List α) : l[l.length]? = none := getElem?_len_le le_rfl +theorem getElem?_length (l : List α) : l[l.length]? = none := getElem?_eq_none le_rfl @[deprecated getElem?_length (since := "2024-06-12")] -theorem get?_length (l : List α) : l.get? l.length = none := get?_len_le le_rfl +theorem get?_length (l : List α) : l.get? l.length = none := get?_eq_none le_rfl @[deprecated (since := "2024-05-03")] alias get?_injective := get?_inj @@ -847,33 +847,11 @@ theorem eq_cons_of_length_one {l : List α} (h : l.length = 1) : l = [l.get ⟨0 end deprecated -theorem modifyTailIdx_modifyTailIdx {f g : List α → List α} (m : ℕ) : - ∀ (n) (l : List α), - (l.modifyTailIdx f n).modifyTailIdx g (m + n) = - l.modifyTailIdx (fun l => (f l).modifyTailIdx g m) n - | 0, _ => rfl - | _ + 1, [] => rfl - | n + 1, a :: l => congr_arg (List.cons a) (modifyTailIdx_modifyTailIdx m n l) - -@[deprecated (since := "2024-10-21")] -alias modifyNthTail_modifyNthTail := modifyTailIdx_modifyTailIdx - -theorem modifyTailIdx_modifyTailIdx_le {f g : List α → List α} (m n : ℕ) (l : List α) - (h : n ≤ m) : - (l.modifyTailIdx f n).modifyTailIdx g m = - l.modifyTailIdx (fun l => (f l).modifyTailIdx g (m - n)) n := by - rcases Nat.exists_eq_add_of_le h with ⟨m, rfl⟩ - rw [Nat.add_comm, modifyTailIdx_modifyTailIdx, Nat.add_sub_cancel] - @[deprecated (since := "2024-10-21")] alias modifyNthTail_modifyNthTail_le := modifyTailIdx_modifyTailIdx_le -theorem modifyTailIdx_modifyTailIdx_same {f g : List α → List α} (n : ℕ) (l : List α) : - (l.modifyTailIdx f n).modifyTailIdx g n = l.modifyTailIdx (g ∘ f) n := by - rw [modifyTailIdx_modifyTailIdx_le n n l (le_refl n), Nat.sub_self]; rfl - @[deprecated (since := "2024-10-21")] -alias modifyNthTail_modifyNthTail_same := modifyTailIdx_modifyTailIdx_same +alias modifyNthTail_modifyNthTail_same := modifyTailIdx_modifyTailIdx_self @[deprecated (since := "2024-05-04")] alias removeNth_eq_nthTail := eraseIdx_eq_modifyTailIdx @[deprecated (since := "2024-10-21")] alias modifyNth_eq_set := modify_eq_set @@ -1043,8 +1021,8 @@ theorem zipWith_flip (f : α → β → γ) : ∀ as bs, zipWith (flip f) bs as (x.drop m).take n ++ x.drop (n + m) = x.drop m := by rw [Nat.add_comm, drop_take_append_drop] /-- `take_concat_get` in simp normal form -/ -@[simp] lemma take_concat_get' (l : List α) (i : ℕ) (h : i < l.length) : - l.take i ++ [l[i]] = l.take (i + 1) := by simpa using take_concat_get l i h +lemma take_concat_get' (l : List α) (i : ℕ) (h : i < l.length) : + l.take i ++ [l[i]] = l.take (i + 1) := by simp /-- `eq_nil_or_concat` in simp normal form -/ lemma eq_nil_or_concat' (l : List α) : l = [] ∨ ∃ L b, l = L ++ [b] := by diff --git a/Mathlib/Data/List/Cycle.lean b/Mathlib/Data/List/Cycle.lean index 167f0c1116671..34485c3ee467f 100644 --- a/Mathlib/Data/List/Cycle.lean +++ b/Mathlib/Data/List/Cycle.lean @@ -217,7 +217,7 @@ theorem prev_ne_cons_cons (y z : α) (h : x ∈ y :: z :: l) (hy : x ≠ y) (hz · rw [prev, dif_neg hy, if_neg hz] theorem next_mem (h : x ∈ l) : l.next x h ∈ l := - nextOr_mem (get_mem _ _ _) + nextOr_mem (get_mem _ _) theorem prev_mem (h : x ∈ l) : l.prev x h ∈ l := by cases' l with hd tl @@ -233,7 +233,7 @@ theorem prev_mem (h : x ∈ l) : l.prev x h ∈ l := by · exact mem_cons_of_mem _ (hl _ _) theorem next_get (l : List α) (h : Nodup l) (i : Fin l.length) : - next l (l.get i) (get_mem _ _ _) = + next l (l.get i) (get_mem _ _) = l.get ⟨(i + 1) % l.length, Nat.mod_lt _ (i.1.zero_le.trans_lt i.2)⟩ := match l, h, i with | [], _, i => by simpa using i.2 @@ -254,7 +254,7 @@ theorem next_get (l : List α) (h : Nodup l) (i : Fin l.length) : · subst hi' rw [next_getLast_cons] · simp [hi', get] - · rw [get_cons_succ]; exact get_mem _ _ _ + · rw [get_cons_succ]; exact get_mem _ _ · exact hx' · simp [getLast_eq_getElem] · exact hn.of_cons @@ -271,12 +271,12 @@ theorem next_get (l : List α) (h : Nodup l) (i : Fin l.length) : intro h have := nodup_iff_injective_get.1 hn h simp at this; simp [this] at hi' - · rw [get_cons_succ]; exact get_mem _ _ _ + · rw [get_cons_succ]; exact get_mem _ _ -- Unused variable linter incorrectly reports that `h` is unused here. set_option linter.unusedVariables false in theorem prev_get (l : List α) (h : Nodup l) (i : Fin l.length) : - prev l (l.get i) (get_mem _ _ _) = + prev l (l.get i) (get_mem _ _) = l.get ⟨(i + (l.length - 1)) % l.length, Nat.mod_lt _ i.pos⟩ := match l with | [] => by simpa using i.2 diff --git a/Mathlib/Data/List/FinRange.lean b/Mathlib/Data/List/FinRange.lean index 83ccab9a12979..8281f5ce92052 100644 --- a/Mathlib/Data/List/FinRange.lean +++ b/Mathlib/Data/List/FinRange.lean @@ -5,7 +5,6 @@ Authors: Mario Carneiro, Kenny Lau, Kim Morrison, Alex Keizer -/ import Mathlib.Data.List.OfFn import Batteries.Data.List.Perm -import Batteries.Data.List.FinRange import Mathlib.Data.List.Nodup /-! diff --git a/Mathlib/Data/List/GetD.lean b/Mathlib/Data/List/GetD.lean index 33af5b90cdb00..5aaf7611ee1bd 100644 --- a/Mathlib/Data/List/GetD.lean +++ b/Mathlib/Data/List/GetD.lean @@ -86,7 +86,7 @@ theorem getD_append_right (l l' : List α) (d : α) (n : ℕ) (h : l.length ≤ theorem getD_eq_getD_get? (n : ℕ) : l.getD n d = (l.get? n).getD d := by cases Nat.lt_or_ge n l.length with | inl h => rw [getD_eq_getElem _ _ h, get?_eq_get h, get_eq_getElem, Option.getD_some] - | inr h => rw [getD_eq_default _ _ h, get?_eq_none.mpr h, Option.getD_none] + | inr h => rw [getD_eq_default _ _ h, get?_eq_none_iff.mpr h, Option.getD_none] end getD diff --git a/Mathlib/Data/List/Indexes.lean b/Mathlib/Data/List/Indexes.lean index fdbe9168bfa47..af5d70ec012f2 100644 --- a/Mathlib/Data/List/Indexes.lean +++ b/Mathlib/Data/List/Indexes.lean @@ -3,7 +3,6 @@ Copyright (c) 2020 Jannis Limperg. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jannis Limperg -/ -import Mathlib.Data.List.OfFn import Mathlib.Data.List.Basic /-! @@ -36,16 +35,11 @@ theorem list_reverse_induction (p : List α → Prop) (base : p []) · apply pq; simp only [reverse_nil, base] · apply pq; simp only [reverse_cons]; apply ind; apply qp; rw [reverse_reverse]; exact ih -@[deprecated (since := "2024-10-15")] alias mapIdxGo_append := mapIdx_go_append @[deprecated (since := "2024-10-15")] alias mapIdxGo_length := mapIdx_go_length -theorem mapIdx_append_one : ∀ (f : ℕ → α → β) (l : List α) (e : α), - mapIdx f (l ++ [e]) = mapIdx f l ++ [f l.length e] := by - intros f l e - unfold mapIdx - rw [mapIdx_go_append] - simp only [mapIdx.go, Array.size_toArray, mapIdx_go_length, length_nil, Nat.add_zero, - Array.push_toList, Array.toList_toArray] +theorem mapIdx_append_one : ∀ {f : ℕ → α → β} {l : List α} {e : α}, + mapIdx f (l ++ [e]) = mapIdx f l ++ [f l.length e] := + mapIdx_concat @[local simp] theorem map_enumFrom_eq_zipWith : ∀ (l : List α) (n : ℕ) (f : ℕ → α → β), diff --git a/Mathlib/Data/List/InsertIdx.lean b/Mathlib/Data/List/InsertIdx.lean index 7c1c61e68cdc2..431b9703de66b 100644 --- a/Mathlib/Data/List/InsertIdx.lean +++ b/Mathlib/Data/List/InsertIdx.lean @@ -27,143 +27,27 @@ section InsertIdx variable {a : α} -@[simp] -theorem insertIdx_zero (s : List α) (x : α) : insertIdx 0 x s = x :: s := - rfl - -@[simp] -theorem insertIdx_succ_nil (n : ℕ) (a : α) : insertIdx (n + 1) a [] = [] := - rfl - -@[simp] -theorem insertIdx_succ_cons (s : List α) (hd x : α) (n : ℕ) : - insertIdx (n + 1) x (hd :: s) = hd :: insertIdx n x s := - rfl - -theorem length_insertIdx : ∀ n as, n ≤ length as → length (insertIdx n a as) = length as + 1 - | 0, _, _ => rfl - | _ + 1, [], h => (Nat.not_succ_le_zero _ h).elim - | n + 1, _ :: as, h => congr_arg Nat.succ <| length_insertIdx n as (Nat.le_of_succ_le_succ h) - -theorem eraseIdx_insertIdx (n : ℕ) (l : List α) : (l.insertIdx n a).eraseIdx n = l := by - rw [eraseIdx_eq_modifyTailIdx, insertIdx, modifyTailIdx_modifyTailIdx_same] - exact modifyTailIdx_id _ _ - -theorem insertIdx_eraseIdx_of_ge : - ∀ n m as, - n < length as → n ≤ m → insertIdx m a (as.eraseIdx n) = (as.insertIdx (m + 1) a).eraseIdx n - | 0, 0, [], has, _ => (lt_irrefl _ has).elim - | 0, 0, _ :: as, _, _ => by simp [eraseIdx, insertIdx] - | 0, _ + 1, _ :: _, _, _ => rfl - | n + 1, m + 1, a :: as, has, hmn => - congr_arg (cons a) <| - insertIdx_eraseIdx_of_ge n m as (Nat.lt_of_succ_lt_succ has) (Nat.le_of_succ_le_succ hmn) - -theorem insertIdx_eraseIdx_of_le : - ∀ n m as, - n < length as → m ≤ n → insertIdx m a (as.eraseIdx n) = (as.insertIdx m a).eraseIdx (n + 1) - | _, 0, _ :: _, _, _ => rfl - | n + 1, m + 1, a :: as, has, hmn => - congr_arg (cons a) <| - insertIdx_eraseIdx_of_le n m as (Nat.lt_of_succ_lt_succ has) (Nat.le_of_succ_le_succ hmn) - -theorem insertIdx_comm (a b : α) : - ∀ (i j : ℕ) (l : List α) (_ : i ≤ j) (_ : j ≤ length l), - (l.insertIdx i a).insertIdx (j + 1) b = (l.insertIdx j b).insertIdx i a - | 0, j, l => by simp [insertIdx] - | _ + 1, 0, _ => fun h => (Nat.not_lt_zero _ h).elim - | i + 1, j + 1, [] => by simp - | i + 1, j + 1, c :: l => fun h₀ h₁ => by - simp only [insertIdx_succ_cons, cons.injEq, true_and] - exact insertIdx_comm a b i j l (Nat.le_of_succ_le_succ h₀) (Nat.le_of_succ_le_succ h₁) - -theorem mem_insertIdx {a b : α} : - ∀ {n : ℕ} {l : List α} (_ : n ≤ l.length), a ∈ l.insertIdx n b ↔ a = b ∨ a ∈ l - | 0, as, _ => by simp - | _ + 1, [], h => (Nat.not_succ_le_zero _ h).elim - | n + 1, a' :: as, h => by - rw [List.insertIdx_succ_cons, mem_cons, mem_insertIdx (Nat.le_of_succ_le_succ h), - ← or_assoc, @or_comm (a = a'), or_assoc, mem_cons] - -theorem insertIdx_of_length_lt (l : List α) (x : α) (n : ℕ) (h : l.length < n) : - insertIdx n x l = l := by - induction' l with hd tl IH generalizing n - · cases n - · simp at h - · simp - · cases n - · simp at h - · simp only [Nat.succ_lt_succ_iff, length] at h - simpa using IH _ h - -@[simp] -theorem insertIdx_length_self (l : List α) (x : α) : insertIdx l.length x l = l ++ [x] := by - induction' l with hd tl IH - · simp - · simpa using IH - -theorem length_le_length_insertIdx (l : List α) (x : α) (n : ℕ) : - l.length ≤ (insertIdx n x l).length := by - rcases le_or_lt n l.length with hn | hn - · rw [length_insertIdx _ _ hn] - exact (Nat.lt_succ_self _).le - · rw [insertIdx_of_length_lt _ _ _ hn] - -theorem length_insertIdx_le_succ (l : List α) (x : α) (n : ℕ) : - (insertIdx n x l).length ≤ l.length + 1 := by - rcases le_or_lt n l.length with hn | hn - · rw [length_insertIdx _ _ hn] - · rw [insertIdx_of_length_lt _ _ _ hn] - exact (Nat.lt_succ_self _).le - -theorem getElem_insertIdx_of_lt (l : List α) (x : α) (n k : ℕ) (hn : k < n) (hk : k < l.length) - (hk' : k < (insertIdx n x l).length := hk.trans_le (length_le_length_insertIdx _ _ _)) : - (insertIdx n x l)[k] = l[k] := by - induction' n with n IH generalizing k l - · simp at hn - · cases' l with hd tl - · simp - · cases k - · simp [get] - · rw [Nat.succ_lt_succ_iff] at hn - simpa using IH _ _ hn _ - theorem get_insertIdx_of_lt (l : List α) (x : α) (n k : ℕ) (hn : k < n) (hk : k < l.length) (hk' : k < (insertIdx n x l).length := hk.trans_le (length_le_length_insertIdx _ _ _)) : (insertIdx n x l).get ⟨k, hk'⟩ = l.get ⟨k, hk⟩ := by simp_all [getElem_insertIdx_of_lt] -@[simp] -theorem getElem_insertIdx_self (l : List α) (x : α) (n : ℕ) (hn : n ≤ l.length) - (hn' : n < (insertIdx n x l).length := (by rwa [length_insertIdx _ _ hn, Nat.lt_succ_iff])) : - (insertIdx n x l)[n] = x := by - induction' l with hd tl IH generalizing n - · simp only [length] at hn - cases hn - simp only [insertIdx_zero, getElem_singleton] - · cases n - · simp - · simp only [Nat.succ_le_succ_iff, length] at hn - simpa using IH _ hn - theorem get_insertIdx_self (l : List α) (x : α) (n : ℕ) (hn : n ≤ l.length) - (hn' : n < (insertIdx n x l).length := (by rwa [length_insertIdx _ _ hn, Nat.lt_succ_iff])) : + (hn' : n < (insertIdx n x l).length := + (by rwa [length_insertIdx_of_le_length hn, Nat.lt_succ_iff])) : (insertIdx n x l).get ⟨n, hn'⟩ = x := by simp [hn, hn'] theorem getElem_insertIdx_add_succ (l : List α) (x : α) (n k : ℕ) (hk' : n + k < l.length) (hk : n + k + 1 < (insertIdx n x l).length := (by - rwa [length_insertIdx _ _ (by omega), Nat.succ_lt_succ_iff])) : + rwa [length_insertIdx_of_le_length (by omega), Nat.succ_lt_succ_iff])) : (insertIdx n x l)[n + k + 1] = l[n + k] := by - induction' l with hd tl IH generalizing n k - · simp at hk' - · cases n - · simp - · simpa [Nat.add_right_comm] using IH _ _ _ + rw [getElem_insertIdx_of_ge (by omega)] + simp only [Nat.add_one_sub_one] theorem get_insertIdx_add_succ (l : List α) (x : α) (n k : ℕ) (hk' : n + k < l.length) (hk : n + k + 1 < (insertIdx n x l).length := (by - rwa [length_insertIdx _ _ (by omega), Nat.succ_lt_succ_iff])) : + rwa [length_insertIdx_of_le_length (by omega), Nat.succ_lt_succ_iff])) : (insertIdx n x l).get ⟨n + k + 1, hk⟩ = get l ⟨n + k, hk'⟩ := by simp [getElem_insertIdx_add_succ, hk, hk'] diff --git a/Mathlib/Data/List/Intervals.lean b/Mathlib/Data/List/Intervals.lean index 9c89237d400af..67e0ed0acd83e 100644 --- a/Mathlib/Data/List/Intervals.lean +++ b/Mathlib/Data/List/Intervals.lean @@ -130,7 +130,7 @@ theorem not_mem_top {n m : ℕ} : m ∉ Ico n m := by simp theorem filter_lt_of_top_le {n m l : ℕ} (hml : m ≤ l) : ((Ico n m).filter fun x => x < l) = Ico n m := filter_eq_self.2 fun k hk => by - simp only [(lt_of_lt_of_le (mem.1 hk).2 hml), decide_True] + 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_iff.2 fun k hk => by diff --git a/Mathlib/Data/List/NodupEquivFin.lean b/Mathlib/Data/List/NodupEquivFin.lean index 65a48a57e4574..b53d6dbaf8777 100644 --- a/Mathlib/Data/List/NodupEquivFin.lean +++ b/Mathlib/Data/List/NodupEquivFin.lean @@ -49,7 +49,7 @@ variable [DecidableEq α] the set of elements of `l`. -/ @[simps] def getEquiv (l : List α) (H : Nodup l) : Fin (length l) ≃ { x // x ∈ l } where - toFun i := ⟨get l i, get_mem l i i.2⟩ + toFun i := ⟨get l i, get_mem _ _⟩ invFun x := ⟨indexOf (↑x) l, indexOf_lt_length.2 x.2⟩ left_inv i := by simp only [List.get_indexOf, eq_self_iff_true, Fin.eta, Subtype.coe_mk, H] right_inv x := by simp @@ -108,7 +108,7 @@ theorem sublist_of_orderEmbedding_get?_eq {l l' : List α} (f : ℕ ↪o ℕ) induction' l with hd tl IH generalizing l' f · simp have : some hd = _ := hf 0 - rw [eq_comm, List.get?_eq_some] at this + rw [eq_comm, List.get?_eq_some_iff] at this obtain ⟨w, h⟩ := this let f' : ℕ ↪o ℕ := OrderEmbedding.ofMapLEIff (fun i => f (i + 1) - (f 0 + 1)) fun a b => by @@ -168,7 +168,7 @@ theorem sublist_iff_exists_fin_orderEmbedding_get_eq {l l' : List α} : have h : ∀ {i : ℕ}, i < l.length → f i < l'.length := by intro i hi specialize hf i - rw [get?_eq_get hi, eq_comm, get?_eq_some] at hf + rw [get?_eq_get hi, eq_comm, get?_eq_some_iff] at hf obtain ⟨h, -⟩ := hf exact h refine ⟨OrderEmbedding.ofMapLEIff (fun ix => ⟨f ix, h ix.is_lt⟩) ?_, ?_⟩ @@ -193,7 +193,7 @@ theorem sublist_iff_exists_fin_orderEmbedding_get_eq {l l' : List α} : simp only [OrderEmbedding.coe_ofStrictMono] split_ifs with hi · rw [get?_eq_get hi, get?_eq_get, ← hf] - · rw [get?_eq_none.mpr, get?_eq_none.mpr] + · rw [get?_eq_none_iff.mpr, get?_eq_none_iff.mpr] · simp · simpa using hi diff --git a/Mathlib/Data/List/OfFn.lean b/Mathlib/Data/List/OfFn.lean index 0e55bda4eaa55..ac14755a5fd6d 100644 --- a/Mathlib/Data/List/OfFn.lean +++ b/Mathlib/Data/List/OfFn.lean @@ -3,7 +3,6 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Batteries.Data.List.OfFn import Mathlib.Data.Fin.Tuple.Basic /-! @@ -36,7 +35,7 @@ theorem get_ofFn {n} (f : Fin n → α) (i) : get (ofFn f) i = f (Fin.cast (by s /-- The `n`th element of a list -/ theorem get?_ofFn {n} (f : Fin n → α) (i) : get? (ofFn f) i = ofFnNthVal f i := by - simp + simp [ofFnNthVal] @[simp] theorem map_ofFn {β : Type*} {n : ℕ} (f : Fin n → α) (g : α → β) : @@ -59,18 +58,6 @@ theorem ofFn_congr {m n : ℕ} (h : m = n) (f : Fin m → α) : subst h simp_rw [Fin.cast_refl, id] -/-- `ofFn` on an empty domain is the empty list. -/ -@[simp] -theorem ofFn_zero (f : Fin 0 → α) : ofFn f = [] := - ext_get (by simp) (fun i hi₁ hi₂ => by contradiction) - -@[simp] -theorem ofFn_succ {n} (f : Fin (succ n) → α) : ofFn f = f 0 :: ofFn fun i => f i.succ := - ext_get (by simp) (fun i hi₁ hi₂ => by - cases i - · simp - · simp) - theorem ofFn_succ' {n} (f : Fin (succ n) → α) : ofFn f = (ofFn fun i => f (Fin.castSucc i)).concat (f (Fin.last _)) := by induction' n with n IH @@ -79,10 +66,6 @@ theorem ofFn_succ' {n} (f : Fin (succ n) → α) : · rw [ofFn_succ, IH, ofFn_succ, concat_cons, Fin.castSucc_zero] congr -@[simp] -theorem ofFn_eq_nil_iff {n : ℕ} {f : Fin n → α} : ofFn f = [] ↔ n = 0 := by - cases n <;> simp only [ofFn_zero, ofFn_succ, eq_self_iff_true, Nat.succ_ne_zero, reduceCtorEq] - /-- Note this matches the convention of `List.ofFn_succ'`, putting the `Fin m` elements first. -/ theorem ofFn_add {m n} (f : Fin (m + n) → α) : List.ofFn f = @@ -172,15 +155,6 @@ theorem pairwise_ofFn {R : α → α → Prop} {n} {f : Fin n → α} : (Fin.rightInverse_cast (length_ofFn f)).surjective.forall, Fin.forall_iff, Fin.cast_mk, Fin.mk_lt_mk, forall_comm (α := (_ : Prop)) (β := ℕ)] -lemma head_ofFn {n} (f : Fin n → α) (h : ofFn f ≠ []) : - (ofFn f).head h = f ⟨0, Nat.pos_of_ne_zero (mt ofFn_eq_nil_iff.2 h)⟩ := by - rw [← getElem_zero (length_ofFn _ ▸ Nat.pos_of_ne_zero (mt ofFn_eq_nil_iff.2 h)), - List.getElem_ofFn] - -lemma getLast_ofFn {n} (f : Fin n → α) (h : ofFn f ≠ []) : - (ofFn f).getLast h = f ⟨n - 1, Nat.sub_one_lt (mt ofFn_eq_nil_iff.2 h)⟩ := by - simp [getLast_eq_getElem] - lemma getLast_ofFn_succ {n : ℕ} (f : Fin n.succ → α) : (ofFn f).getLast (mt ofFn_eq_nil_iff.1 (Nat.succ_ne_zero _)) = f (Fin.last _) := getLast_ofFn f _ @@ -200,31 +174,6 @@ lemma ofFn_cons {n} (a : α) (f : Fin n → α) : ofFn (Fin.cons a f) = a :: ofF rw [ofFn_succ] rfl --- Temporary local copy of result from Lean commit 1e98fd7f2d965ab035dbf1099fb4a4ffde16b151. -theorem find?_eq_some_iff_getElem {xs : List α} {p : α → Bool} {b : α} : - xs.find? p = some b ↔ p b ∧ ∃ i h, xs[i] = b ∧ ∀ j : Nat, (hj : j < i) → !p xs[j] := by - rw [find?_eq_some] - simp only [Bool.not_eq_eq_eq_not, Bool.not_true, exists_and_right, and_congr_right_iff] - intro w - constructor - · rintro ⟨as, ⟨bs, rfl⟩, h⟩ - refine ⟨as.length, ⟨?_, ?_, ?_⟩⟩ - · simp only [length_append, length_cons] - refine Nat.lt_add_of_pos_right (zero_lt_succ bs.length) - · rw [getElem_append_right (Nat.le_refl as.length)] - simp - · intro j h' - rw [getElem_append_left h'] - exact h _ (getElem_mem h') - · rintro ⟨i, h, rfl, h'⟩ - refine ⟨xs.take i, ⟨xs.drop (i+1), ?_⟩, ?_⟩ - · rw [getElem_cons_drop, take_append_drop] - · intro a m - rw [mem_take_iff_getElem] at m - obtain ⟨j, h, rfl⟩ := m - apply h' - omega - lemma find?_ofFn_eq_some {n} {f : Fin n → α} {p : α → Bool} {b : α} : (ofFn f).find? p = some b ↔ p b = true ∧ ∃ i, f i = b ∧ ∀ j < i, ¬(p (f j) = true) := by rw [find?_eq_some_iff_getElem] diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index 3e0073e3e2049..f436730690267 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -410,7 +410,7 @@ theorem count_permutations'Aux_self [DecidableEq α] (l : List α) (x : α) : simpa [takeWhile, Nat.succ_inj', DecEq_eq] using IH _ · rw [takeWhile] simp only [mem_map, cons.injEq, Ne.symm hx, false_and, and_false, exists_false, - not_false_iff, count_eq_zero_of_not_mem, Nat.zero_add, hx, decide_False, length_nil] + not_false_iff, count_eq_zero_of_not_mem, Nat.zero_add, hx, decide_false, length_nil] @[simp] theorem length_permutations'Aux (s : List α) (x : α) : @@ -454,7 +454,7 @@ theorem nodup_permutations'Aux_iff {s : List α} {x : α} : Nodup (permutations' rw [← @Fin.mk.inj_iff _ _ _ kl k1l]; apply h rw [get_permutations'Aux, get_permutations'Aux] have hl : length (insertIdx k x s) = length (insertIdx (k + 1) x s) := by - rw [length_insertIdx _ _ hk.le, length_insertIdx _ _ (Nat.succ_le_of_lt hk)] + rw [length_insertIdx_of_le_length hk.le, length_insertIdx_of_le_length (Nat.succ_le_of_lt hk)] refine ext_get hl fun n hn hn' => ?_ rcases lt_trichotomy n k with (H | rfl | H) · rw [get_insertIdx_of_lt _ _ _ _ H (H.trans hk), @@ -465,7 +465,7 @@ theorem nodup_permutations'Aux_iff {s : List α} {x : α} : Nodup (permutations' convert hk' using 1 exact get_insertIdx_add_succ _ _ _ 0 _ · obtain ⟨m, rfl⟩ := Nat.exists_eq_add_of_lt H' - rw [length_insertIdx _ _ hk.le, Nat.succ_lt_succ_iff, Nat.succ_add] at hn + rw [length_insertIdx_of_le_length hk.le, Nat.succ_lt_succ_iff, Nat.succ_add] at hn rw [get_insertIdx_add_succ] · convert get_insertIdx_add_succ s x k m.succ (by simpa using hn) using 2 · simp [Nat.add_assoc, Nat.add_left_comm] @@ -492,21 +492,21 @@ theorem nodup_permutations (s : List α) (hs : Nodup s) : Nodup s.permutations : have hl : as.length = bs.length := (ha.trans hb.symm).length_eq simp only [Nat.lt_succ_iff, length_permutations'Aux] at hn hm rw [get_permutations'Aux] at hn' hm' - have hx : - (insertIdx n x as)[m]'(by rwa [length_insertIdx _ _ hn, Nat.lt_succ_iff, hl]) = x := by + have hx : (insertIdx n x as)[m]'(by + rwa [length_insertIdx_of_le_length hn, Nat.lt_succ_iff, hl]) = x := by simp [hn', ← hm', hm] - have hx' : - (insertIdx m x bs)[n]'(by rwa [length_insertIdx _ _ hm, Nat.lt_succ_iff, ← hl]) = x := by + have hx' : (insertIdx m x bs)[n]'(by + rwa [length_insertIdx_of_le_length hm, Nat.lt_succ_iff, ← hl]) = x := by simp [hm', ← hn', hn] rcases lt_trichotomy n m with (ht | ht | ht) · suffices x ∈ bs by exact h x (hb.subset this) rfl - rw [← hx', getElem_insertIdx_of_lt _ _ _ _ ht (ht.trans_le hm)] + rw [← hx', getElem_insertIdx_of_lt ht] exact getElem_mem _ · simp only [ht] at hm' hn' rw [← hm'] at hn' exact H (insertIdx_injective _ _ hn') · suffices x ∈ as by exact h x (ha.subset this) rfl - rw [← hx, getElem_insertIdx_of_lt _ _ _ _ ht (ht.trans_le hn)] + rw [← hx, getElem_insertIdx_of_lt ht] exact getElem_mem _ lemma permutations_take_two (x y : α) (s : List α) : diff --git a/Mathlib/Data/List/Sublists.lean b/Mathlib/Data/List/Sublists.lean index 9a2c8391406a8..aca81a62467f5 100644 --- a/Mathlib/Data/List/Sublists.lean +++ b/Mathlib/Data/List/Sublists.lean @@ -45,7 +45,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_toList] + rw [sublists'Aux, Array.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 @@ -107,7 +107,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_toList, Array.mkEmpty] + simp only [sublistsAux, Array.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) @@ -128,7 +128,7 @@ theorem sublistsAux_eq_flatMap : ext α l : 2 trans l.foldr sublistsAux [[]] · rw [sublistsAux_eq_flatMap, sublists] - · simp only [sublistsFast, sublistsAux_eq_array_foldl, Array.foldr_eq_foldr_toList] + · simp only [sublistsFast, sublistsAux_eq_array_foldl, Array.foldr_toList] rw [← foldr_hom Array.toList] · intros _ _; congr diff --git a/Mathlib/Data/List/Sym.lean b/Mathlib/Data/List/Sym.lean index 1d6228b9258e3..b3981a8ba0333 100644 --- a/Mathlib/Data/List/Sym.lean +++ b/Mathlib/Data/List/Sym.lean @@ -170,7 +170,7 @@ theorem dedup_sym2 [DecidableEq α] (xs : List α) : xs.sym2.dedup = xs.dedup.sy obtain hm | hm := Decidable.em (x ∈ xs) · rw [dedup_cons_of_mem hm, ← ih, dedup_cons_of_mem, List.Subset.dedup_append_right (map_mk_sublist_sym2 _ _ hm).subset] - refine mem_append_of_mem_left _ ?_ + refine mem_append_left _ ?_ rw [mem_map] exact ⟨_, hm, Sym2.eq_swap⟩ · rw [dedup_cons_of_not_mem hm, List.sym2, map_cons, ← ih, dedup_cons_of_not_mem, cons_append, diff --git a/Mathlib/Data/List/TFAE.lean b/Mathlib/Data/List/TFAE.lean index 24e43d9612f03..55f77a0f094f9 100644 --- a/Mathlib/Data/List/TFAE.lean +++ b/Mathlib/Data/List/TFAE.lean @@ -62,7 +62,7 @@ theorem tfae_of_cycle {a b} {l : List Prop} (h_chain : List.Chain (· → ·) a theorem TFAE.out {l} (h : TFAE l) (n₁ n₂) {a b} (h₁ : List.get? l n₁ = some a := by rfl) (h₂ : List.get? l n₂ = some b := by rfl) : a ↔ b := - h _ (List.get?_mem h₁) _ (List.get?_mem h₂) + h _ (List.mem_of_get? h₁) _ (List.mem_of_get? h₂) /-- If `P₁ x ↔ ... ↔ Pₙ x` for all `x`, then `(∀ x, P₁ x) ↔ ... ↔ (∀ x, Pₙ x)`. Note: in concrete cases, Lean has trouble finding the list `[P₁, ..., Pₙ]` from the list diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index e364da53cefec..83aedd9344065 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -1856,7 +1856,7 @@ theorem filter_add_not (s : Multiset α) : filter p s + filter (fun a => ¬p a) · simp only [add_zero] · simp [Decidable.em, -Bool.not_eq_true, -not_and, not_and_or, or_comm] · simp only [Bool.not_eq_true, decide_eq_true_eq, Bool.eq_false_or_eq_true, - decide_True, implies_true, Decidable.em] + decide_true, implies_true, Decidable.em] theorem filter_map (f : β → α) (s : Multiset β) : filter p (map f s) = map f (filter (p ∘ f) s) := Quot.inductionOn s fun l => by simp [List.filter_map]; rfl diff --git a/Mathlib/Data/NNReal/Defs.lean b/Mathlib/Data/NNReal/Defs.lean index 9c117c2ed6f81..c361f32478764 100644 --- a/Mathlib/Data/NNReal/Defs.lean +++ b/Mathlib/Data/NNReal/Defs.lean @@ -53,7 +53,7 @@ open Function /-- Nonnegative real numbers. -/ def NNReal := { r : ℝ // 0 ≤ r } deriving Zero, One, Semiring, StrictOrderedSemiring, CommMonoidWithZero, CommSemiring, - SemilatticeInf, SemilatticeSup, DistribLattice, OrderedCommSemiring, + PartialOrder, SemilatticeInf, SemilatticeSup, DistribLattice, OrderedCommSemiring, CanonicallyOrderedCommSemiring, Inhabited namespace NNReal diff --git a/Mathlib/Data/Nat/BitIndices.lean b/Mathlib/Data/Nat/BitIndices.lean index 9d9c7c094edad..8c4e3d74b89f5 100644 --- a/Mathlib/Data/Nat/BitIndices.lean +++ b/Mathlib/Data/Nat/BitIndices.lean @@ -114,6 +114,6 @@ theorem two_pow_le_of_mem_bitIndices (ha : a ∈ n.bitIndices) : 2^a ≤ n := by exact List.single_le_sum (by simp) _ <| mem_map_of_mem _ ha theorem not_mem_bitIndices_self (n : ℕ) : n ∉ n.bitIndices := - fun h ↦ (lt_two_pow n).not_le <| two_pow_le_of_mem_bitIndices h + fun h ↦ (n.lt_two_pow_self).not_le <| two_pow_le_of_mem_bitIndices h end Nat diff --git a/Mathlib/Data/Nat/Bitwise.lean b/Mathlib/Data/Nat/Bitwise.lean index e3b32b271b950..cf9e742f04bfc 100644 --- a/Mathlib/Data/Nat/Bitwise.lean +++ b/Mathlib/Data/Nat/Bitwise.lean @@ -51,7 +51,7 @@ lemma bitwise_zero_left (m : Nat) : bitwise f 0 m = if f false true then m else @[simp] lemma bitwise_zero_right (n : Nat) : bitwise f n 0 = if f true false then n else 0 := by unfold bitwise - simp only [ite_self, decide_False, Nat.zero_div, ite_true, ite_eq_right_iff] + simp only [ite_self, decide_false, Nat.zero_div, ite_true, ite_eq_right_iff] rintro ⟨⟩ split_ifs <;> rfl diff --git a/Mathlib/Data/Nat/Choose/Lucas.lean b/Mathlib/Data/Nat/Choose/Lucas.lean index e3a167ba89664..4e9f92cd773fa 100644 --- a/Mathlib/Data/Nat/Choose/Lucas.lean +++ b/Mathlib/Data/Nat/Choose/Lucas.lean @@ -48,8 +48,8 @@ theorem choose_modEq_choose_mod_mul_choose_div : rw [Prod.mk.injEq] constructor <;> intro h · simp only [mem_product, mem_range] at hx - have h' : x₁ < p := lt_of_lt_of_le hx.left <| mod_lt _ Fin.size_pos' - rw [h, add_mul_mod_self_left, add_mul_div_left _ _ Fin.size_pos', eq_comm (b := x₂)] + have h' : x₁ < p := lt_of_lt_of_le hx.left <| mod_lt _ Fin.pos' + rw [h, add_mul_mod_self_left, add_mul_div_left _ _ Fin.pos', eq_comm (b := x₂)] exact ⟨mod_eq_of_lt h', self_eq_add_left.mpr (div_eq_of_lt h')⟩ · rw [← h.left, ← h.right, mod_add_div] simp only [finset_sum_coeff, coeff_mul_natCast, coeff_X_pow, ite_mul, zero_mul, ← cast_mul] diff --git a/Mathlib/Data/Nat/Count.lean b/Mathlib/Data/Nat/Count.lean index 4d8294ce6bceb..0211381c6106a 100644 --- a/Mathlib/Data/Nat/Count.lean +++ b/Mathlib/Data/Nat/Count.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Vladimir Goryachev, Kyle Miller, Kim Morrison, Eric Rodriguez -/ import Mathlib.SetTheory.Cardinal.Basic +import Mathlib.Algebra.Order.Ring.Nat /-! # Counting on ℕ diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index 79583654cc19e..fc518c6dbbc1c 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -622,15 +622,6 @@ lemma le_self_pow (hn : n ≠ 0) : ∀ a : ℕ, a ≤ a ^ n | 0 => zero_le _ | a + 1 => by simpa using Nat.pow_le_pow_right a.succ_pos (Nat.one_le_iff_ne_zero.2 hn) -lemma lt_pow_self (ha : 1 < a) : ∀ n : ℕ, n < a ^ n - | 0 => by simp - | n + 1 => - calc - n + 1 < a ^ n + 1 := Nat.add_lt_add_right (lt_pow_self ha _) _ - _ ≤ a ^ (n + 1) := Nat.pow_lt_pow_succ ha - -lemma lt_two_pow (n : ℕ) : n < 2 ^ n := lt_pow_self (by decide) n - lemma one_le_pow (n m : ℕ) (h : 0 < m) : 1 ≤ m ^ n := by simpa using Nat.pow_le_pow_of_le_left h n lemma one_le_pow' (n m : ℕ) : 1 ≤ (m + 1) ^ n := one_le_pow n (m + 1) (succ_pos m) @@ -653,7 +644,7 @@ lemma one_lt_two_pow' (n : ℕ) : 1 < 2 ^ (n + 1) := one_lt_pow n.succ_ne_zero ( lemma mul_lt_mul_pow_succ (ha : 0 < a) (hb : 1 < b) : n * b < a * b ^ (n + 1) := by rw [Nat.pow_succ, ← Nat.mul_assoc, Nat.mul_lt_mul_right (Nat.lt_trans Nat.zero_lt_one hb)] exact Nat.lt_of_le_of_lt (Nat.le_mul_of_pos_left _ ha) - ((Nat.mul_lt_mul_left ha).2 <| Nat.lt_pow_self hb _) + ((Nat.mul_lt_mul_left ha).2 <| Nat.lt_pow_self hb) lemma sq_sub_sq (a b : ℕ) : a ^ 2 - b ^ 2 = (a + b) * (a - b) := by simpa [Nat.pow_succ] using Nat.mul_self_sub_mul_self_eq a b @@ -1033,7 +1024,7 @@ lemma not_pos_pow_dvd : ∀ {a n : ℕ} (_ : 1 < a) (_ : 1 < n), ¬ a ^ n ∣ a have : succ a * succ a ^ n ∣ succ a * 1 := by simpa [pow_succ'] using h have : succ a ^ n ∣ 1 := Nat.dvd_of_mul_dvd_mul_left (succ_pos _) this have he : succ a ^ n = 1 := eq_one_of_dvd_one this - have : n < succ a ^ n := lt_pow_self hp n + have : n < succ a ^ n := n.lt_pow_self hp have : n < 1 := by rwa [he] at this have : n = 0 := Nat.eq_zero_of_le_zero <| le_of_lt_succ this have : 1 < 1 := by rwa [this] at hk @@ -1041,7 +1032,7 @@ lemma not_pos_pow_dvd : ∀ {a n : ℕ} (_ : 1 < a) (_ : 1 < n), ¬ a ^ n ∣ a lemma lt_of_pow_dvd_right (hb : b ≠ 0) (ha : 2 ≤ a) (h : a ^ n ∣ b) : n < b := by rw [← Nat.pow_lt_pow_iff_right (succ_le_iff.1 ha)] - exact Nat.lt_of_le_of_lt (le_of_dvd (Nat.pos_iff_ne_zero.2 hb) h) (lt_pow_self ha _) + exact Nat.lt_of_le_of_lt (le_of_dvd (Nat.pos_iff_ne_zero.2 hb) h) (Nat.lt_pow_self ha) lemma div_dvd_of_dvd (h : n ∣ m) : m / n ∣ m := ⟨n, (Nat.div_mul_cancel h).symm⟩ diff --git a/Mathlib/Data/Nat/Factorization/Basic.lean b/Mathlib/Data/Nat/Factorization/Basic.lean index 4e4bd0fa951c0..15659840ed9b9 100644 --- a/Mathlib/Data/Nat/Factorization/Basic.lean +++ b/Mathlib/Data/Nat/Factorization/Basic.lean @@ -155,7 +155,7 @@ theorem ordCompl_mul (a b p : ℕ) : ordCompl[p] (a * b) = ordCompl[p] a * ordCo theorem factorization_lt {n : ℕ} (p : ℕ) (hn : n ≠ 0) : n.factorization p < n := by by_cases pp : p.Prime · exact (Nat.pow_lt_pow_iff_right pp.one_lt).1 <| (ordProj_le p hn).trans_lt <| - lt_pow_self pp.one_lt _ + Nat.lt_pow_self pp.one_lt · simpa only [factorization_eq_zero_of_non_prime n pp] using hn.bot_lt /-- An upper bound on `n.factorization p` -/ @@ -356,7 +356,7 @@ theorem dvd_iff_prime_pow_dvd_dvd (n d : ℕ) : · simp rcases eq_or_ne d 0 with (rfl | hd) · simp only [zero_dvd_iff, hn, false_iff, not_forall] - exact ⟨2, n, prime_two, dvd_zero _, mt (le_of_dvd hn.bot_lt) (lt_two_pow n).not_le⟩ + exact ⟨2, n, prime_two, dvd_zero _, mt (le_of_dvd hn.bot_lt) (n.lt_two_pow_self).not_le⟩ refine ⟨fun h p k _ hpkd => dvd_trans hpkd h, ?_⟩ rw [← factorization_prime_le_iff_dvd hd hn] intro h p pp diff --git a/Mathlib/Data/Nat/Log.lean b/Mathlib/Data/Nat/Log.lean index 1d16748dda19e..b2e55d35fe599 100644 --- a/Mathlib/Data/Nat/Log.lean +++ b/Mathlib/Data/Nat/Log.lean @@ -123,7 +123,7 @@ theorem lt_pow_of_log_lt {b x y : ℕ} (hb : 1 < b) : log b y < x → y < b ^ x lemma log_lt_self (b : ℕ) {x : ℕ} (hx : x ≠ 0) : log b x < x := match le_or_lt b 1 with | .inl h => log_of_left_le_one h x ▸ Nat.pos_iff_ne_zero.2 hx - | .inr h => log_lt_of_lt_pow hx <| lt_pow_self h _ + | .inr h => log_lt_of_lt_pow hx <| Nat.lt_pow_self h lemma log_le_self (b x : ℕ) : log b x ≤ x := if hx : x = 0 then by simp [hx] diff --git a/Mathlib/Data/Nat/Prime/Basic.lean b/Mathlib/Data/Nat/Prime/Basic.lean index 1802c6c89eab8..68d16200182d3 100644 --- a/Mathlib/Data/Nat/Prime/Basic.lean +++ b/Mathlib/Data/Nat/Prime/Basic.lean @@ -51,14 +51,7 @@ theorem Prime.five_le_of_ne_two_of_ne_three {p : ℕ} (hp : p.Prime) (h_two : p (h_three : p ≠ 3) : 5 ≤ p := by by_contra! h revert h_two h_three hp - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11043): was `decide!` - match p with - | 0 => decide - | 1 => decide - | 2 => decide - | 3 => decide - | 4 => decide - | n + 5 => exact (h.not_le <| le_add_left ..).elim + decide +revert end diff --git a/Mathlib/Data/Num/Bitwise.lean b/Mathlib/Data/Num/Bitwise.lean index b374f02e13946..57703e62a4600 100644 --- a/Mathlib/Data/Num/Bitwise.lean +++ b/Mathlib/Data/Num/Bitwise.lean @@ -409,7 +409,7 @@ end SNum namespace SNum /-- `a.bits n` is the vector of the `n` first bits of `a` (starting from the LSB). -/ -def bits : SNum → ∀ n, Vector Bool n +def bits : SNum → ∀ n, Mathlib.Vector Bool n | _, 0 => Vector.nil | p, n + 1 => head p ::ᵥ bits (tail p) n diff --git a/Mathlib/Data/Rat/Denumerable.lean b/Mathlib/Data/Rat/Denumerable.lean index 9a1494cd76969..63a9a4318fa46 100644 --- a/Mathlib/Data/Rat/Denumerable.lean +++ b/Mathlib/Data/Rat/Denumerable.lean @@ -3,9 +3,9 @@ Copyright (c) 2019 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import Mathlib.Algebra.CharZero.Infinite -import Mathlib.Algebra.Ring.Rat +import Mathlib.Algebra.Order.Ring.Rat import Mathlib.Data.Rat.Encodable +import Mathlib.Algebra.CharZero.Infinite import Mathlib.Logic.Denumerable /-! diff --git a/Mathlib/Data/Seq/Seq.lean b/Mathlib/Data/Seq/Seq.lean index 76e8c653427c3..8b25a1d0c7c0a 100644 --- a/Mathlib/Data/Seq/Seq.lean +++ b/Mathlib/Data/Seq/Seq.lean @@ -376,7 +376,7 @@ theorem coinduction2 (s) (f g : Seq α → Seq β) @[coe] def ofList (l : List α) : Seq α := ⟨List.get? l, fun {n} h => by - rw [List.get?_eq_none] at h ⊢ + rw [List.get?_eq_none_iff] at h ⊢ exact h.trans (Nat.le_succ n)⟩ instance coeList : Coe (List α) (Seq α) := diff --git a/Mathlib/Data/Set/List.lean b/Mathlib/Data/Set/List.lean index 60839e4f416e6..0d3cf309c7743 100644 --- a/Mathlib/Data/Set/List.lean +++ b/Mathlib/Data/Set/List.lean @@ -41,8 +41,10 @@ theorem range_list_get : range l.get = { x | x ∈ l } := by theorem range_list_get? : range l.get? = insert none (some '' { x | x ∈ l }) := by rw [← range_list_get, ← range_comp] refine (range_subset_iff.2 fun n => ?_).antisymm (insert_subset_iff.2 ⟨?_, ?_⟩) - exacts [(le_or_lt l.length n).imp get?_eq_none.2 (fun hlt => ⟨⟨_, hlt⟩, (get?_eq_get hlt).symm⟩), - ⟨_, get?_eq_none.2 le_rfl⟩, range_subset_iff.2 fun k => ⟨_, get?_eq_get _⟩] + · exact (le_or_lt l.length n).imp get?_eq_none_iff.mpr + (fun hlt => ⟨⟨_, hlt⟩, (get?_eq_get hlt).symm⟩) + · exact ⟨_, get?_eq_none_iff.mpr le_rfl⟩ + · exact range_subset_iff.2 fun k => ⟨_, get?_eq_get _⟩ @[simp] theorem range_list_getD (d : α) : (range fun n : Nat => l[n]?.getD d) = insert d { x | x ∈ l } := diff --git a/Mathlib/Data/Sign.lean b/Mathlib/Data/Sign.lean index a2065300b0e34..55941cf1dae35 100644 --- a/Mathlib/Data/Sign.lean +++ b/Mathlib/Data/Sign.lean @@ -122,7 +122,13 @@ instance : BoundedOrder SignType where top := 1 le_top := LE.of_pos bot := -1 - bot_le := LE.of_neg + bot_le := + #adaptation_note + /-- + Added `by exact` after https://github.com/leanprover/lean4/pull/6053, + but don't understand why it was needed. + -/ + by exact LE.of_neg instance : HasDistribNeg SignType := { neg_neg := fun x => by cases x <;> rfl diff --git a/Mathlib/Data/Sym/Basic.lean b/Mathlib/Data/Sym/Basic.lean index ac156160cdcb8..32aca18c1b5c4 100644 --- a/Mathlib/Data/Sym/Basic.lean +++ b/Mathlib/Data/Sym/Basic.lean @@ -55,7 +55,7 @@ instance {α : Type*} {n : ℕ} [DecidableEq α] : DecidableEq (Sym α n) := See note [reducible non-instances]. -/ -abbrev Vector.Perm.isSetoid (α : Type*) (n : ℕ) : Setoid (Vector α n) := +abbrev Mathlib.Vector.Perm.isSetoid (α : Type*) (n : ℕ) : Setoid (Vector α n) := (List.isSetoid α).comap Subtype.val attribute [local instance] Vector.Perm.isSetoid @@ -119,20 +119,21 @@ theorem coe_cons (s : Sym α n) (a : α) : (a ::ₛ s : Multiset α) = a ::ₘ s /-- This is the quotient map that takes a list of n elements as an n-tuple and produces an nth symmetric power. -/ -def ofVector : Vector α n → Sym α n := +def ofVector : Mathlib.Vector α n → Sym α n := fun x => ⟨↑x.val, (Multiset.coe_card _).trans x.2⟩ /-- This is the quotient map that takes a list of n elements as an n-tuple and produces an nth symmetric power. -/ -instance : Coe (Vector α n) (Sym α n) where coe x := ofVector x +instance : Coe (Mathlib.Vector α n) (Sym α n) where coe x := ofVector x @[simp] -theorem ofVector_nil : ↑(Vector.nil : Vector α 0) = (Sym.nil : Sym α 0) := +theorem ofVector_nil : ↑(Vector.nil : Mathlib.Vector α 0) = (Sym.nil : Sym α 0) := rfl @[simp] -theorem ofVector_cons (a : α) (v : Vector α n) : ↑(Vector.cons a v) = a ::ₛ (↑v : Sym α n) := by +theorem ofVector_cons (a : α) (v : Mathlib.Vector α n) : + ↑(Vector.cons a v) = a ::ₛ (↑v : Sym α n) := by cases v rfl @@ -179,13 +180,13 @@ theorem mem_cons_of_mem (h : a ∈ s) : a ∈ b ::ₛ s := theorem mem_cons_self (a : α) (s : Sym α n) : a ∈ a ::ₛ s := Multiset.mem_cons_self a s.1 -theorem cons_of_coe_eq (a : α) (v : Vector α n) : a ::ₛ (↑v : Sym α n) = ↑(a ::ᵥ v) := +theorem cons_of_coe_eq (a : α) (v : Mathlib.Vector α n) : a ::ₛ (↑v : Sym α n) = ↑(a ::ᵥ v) := Subtype.ext <| by cases v rfl open scoped List in -theorem sound {a b : Vector α n} (h : a.val ~ b.val) : (↑a : Sym α n) = ↑b := +theorem sound {a b : Mathlib.Vector α n} (h : a.val ~ b.val) : (↑a : Sym α n) = ↑b := Subtype.ext <| Quotient.sound h /-- `erase s a h` is the sym that subtracts 1 from the diff --git a/Mathlib/Data/Sym/Sym2.lean b/Mathlib/Data/Sym/Sym2.lean index 28fbe8b4427a2..a827017219d56 100644 --- a/Mathlib/Data/Sym/Sym2.lean +++ b/Mathlib/Data/Sym/Sym2.lean @@ -586,7 +586,7 @@ section SymEquiv attribute [local instance] Vector.Perm.isSetoid -private def fromVector : Vector α 2 → α × α +private def fromVector : Mathlib.Vector α 2 → α × α | ⟨[a, b], _⟩ => (a, b) private theorem perm_card_two_iff {a₁ b₁ a₂ b₂ : α} : diff --git a/Mathlib/Data/Vector/Basic.lean b/Mathlib/Data/Vector/Basic.lean index 3739f74c8f540..f965bcd759bc9 100644 --- a/Mathlib/Data/Vector/Basic.lean +++ b/Mathlib/Data/Vector/Basic.lean @@ -7,7 +7,6 @@ import Mathlib.Algebra.BigOperators.Group.List import Mathlib.Data.Vector.Defs import Mathlib.Data.List.Nodup import Mathlib.Data.List.OfFn -import Mathlib.Data.List.InsertIdx import Mathlib.Control.Applicative import Mathlib.Control.Traversable.Basic @@ -527,8 +526,7 @@ variable {a : α} def insertIdx (a : α) (i : Fin (n + 1)) (v : Vector α n) : Vector α (n + 1) := ⟨v.1.insertIdx i a, by rw [List.length_insertIdx, v.2] - rw [v.2, ← Nat.succ_le_succ_iff] - exact i.2⟩ + split <;> omega⟩ @[deprecated (since := "2024-10-21")] alias insertNth := insertIdx diff --git a/Mathlib/Data/Vector/Mem.lean b/Mathlib/Data/Vector/Mem.lean index 3eb0fe6678d82..dc09c5bb58415 100644 --- a/Mathlib/Data/Vector/Mem.lean +++ b/Mathlib/Data/Vector/Mem.lean @@ -22,9 +22,7 @@ namespace Vector variable {α β : Type*} {n : ℕ} (a a' : α) @[simp] -theorem get_mem (i : Fin n) (v : Vector α n) : v.get i ∈ v.toList := by - rw [get_eq_get] - exact List.get_mem _ _ _ +theorem get_mem (i : Fin n) (v : Vector α n) : v.get i ∈ v.toList := List.get_mem _ _ theorem mem_iff_get (v : Vector α n) : a ∈ v.toList ↔ ∃ i, v.get i = a := by simp only [List.mem_iff_get, Fin.exists_iff, Vector.get_eq_get] diff --git a/Mathlib/Deprecated/Order.lean b/Mathlib/Deprecated/Order.lean index 9f79797c7e46f..65fab91526170 100644 --- a/Mathlib/Deprecated/Order.lean +++ b/Mathlib/Deprecated/Order.lean @@ -12,7 +12,7 @@ import Mathlib.Order.Defs.LinearOrder variable {α : Type*} -@[deprecated (since := "2024-11-26")] -- unused in Mathlib +@[deprecated "This was a leftover from Lean 3, and has not been needed." (since := "2024-11-26")] instance isStrictTotalOrder_of_linearOrder [LinearOrder α] : IsStrictTotalOrder α (· < ·) where irrefl := lt_irrefl trichotomous := lt_trichotomy diff --git a/Mathlib/Dynamics/Newton.lean b/Mathlib/Dynamics/Newton.lean index 94030f064065f..3a68a8bf036e7 100644 --- a/Mathlib/Dynamics/Newton.lean +++ b/Mathlib/Dynamics/Newton.lean @@ -111,7 +111,7 @@ theorem exists_unique_nilpotent_sub_and_aeval_eq_zero · -- Existence obtain ⟨n, hn⟩ := id h refine ⟨P.newtonMap^[n] x, isNilpotent_iterate_newtonMap_sub_of_isNilpotent h n, ?_⟩ - rw [← zero_dvd_iff, ← pow_eq_zero_of_le n.lt_two_pow.le hn] + rw [← zero_dvd_iff, ← pow_eq_zero_of_le (n.lt_two_pow_self).le hn] exact aeval_pow_two_pow_dvd_aeval_iterate_newtonMap h h' n · -- Uniqueness have ⟨u, hu⟩ := binomExpansion (P.map (algebraMap R S)) r₁ (r₂ - r₁) diff --git a/Mathlib/FieldTheory/IsPerfectClosure.lean b/Mathlib/FieldTheory/IsPerfectClosure.lean index 5c6d7928dc1a4..db75df108ee8c 100644 --- a/Mathlib/FieldTheory/IsPerfectClosure.lean +++ b/Mathlib/FieldTheory/IsPerfectClosure.lean @@ -96,7 +96,7 @@ theorem mem_pNilradical {R : Type*} [CommSemiring R] {p : ℕ} {x : R} : by_cases hp : 1 < p · rw [pNilradical_eq_nilradical hp] refine ⟨fun ⟨n, h⟩ ↦ ⟨n, ?_⟩, fun ⟨n, h⟩ ↦ ⟨p ^ n, h⟩⟩ - rw [← Nat.sub_add_cancel ((Nat.lt_pow_self hp n).le), pow_add, h, mul_zero] + rw [← Nat.sub_add_cancel ((n.lt_pow_self hp).le), pow_add, h, mul_zero] rw [pNilradical_eq_bot hp, Ideal.mem_bot] refine ⟨fun h ↦ ⟨0, by rw [pow_zero, pow_one, h]⟩, fun ⟨n, h⟩ ↦ ?_⟩ rcases Nat.le_one_iff_eq_zero_or_eq_one.1 (not_lt.1 hp) with hp | hp diff --git a/Mathlib/FieldTheory/KummerExtension.lean b/Mathlib/FieldTheory/KummerExtension.lean index 273c7bd88310d..fee360c7fd681 100644 --- a/Mathlib/FieldTheory/KummerExtension.lean +++ b/Mathlib/FieldTheory/KummerExtension.lean @@ -425,7 +425,7 @@ def adjoinRootXPowSubCEquiv (hζ : (primitiveRoots n K).Nonempty) (H : Irreducib AlgEquiv.ofBijective (AdjoinRoot.liftHom (X ^ n - C a) α (by simp [hα])) <| by haveI := Fact.mk H letI := isSplittingField_AdjoinRoot_X_pow_sub_C hζ H - refine ⟨(AlgHom.toRingHom _).injective, ?_⟩ + refine ⟨(liftHom (X ^ n - C a) α _).injective, ?_⟩ rw [← AlgHom.range_eq_top, ← IsSplittingField.adjoin_rootSet _ (X ^ n - C a), eq_comm, adjoin_rootSet_eq_range, IsSplittingField.adjoin_rootSet] exact IsSplittingField.splits _ _ diff --git a/Mathlib/FieldTheory/PerfectClosure.lean b/Mathlib/FieldTheory/PerfectClosure.lean index 6b6f45e166246..7f553815a9734 100644 --- a/Mathlib/FieldTheory/PerfectClosure.lean +++ b/Mathlib/FieldTheory/PerfectClosure.lean @@ -404,7 +404,7 @@ theorem of_apply (x : K) : of K p x = mk _ _ (0, x) := instance instReduced : IsReduced (PerfectClosure K p) where eq_zero x := induction_on x fun x ⟨n, h⟩ ↦ by replace h : mk K p x ^ p ^ n = 0 := by - rw [← Nat.sub_add_cancel ((Nat.lt_pow_self (Fact.out : p.Prime).one_lt n).le), + rw [← Nat.sub_add_cancel ((n.lt_pow_self (Fact.out : p.Prime).one_lt).le), pow_add, h, mul_zero] simp only [zero_def, mk_pow, mk_eq_iff, zero_add, ← coe_iterateFrobenius, map_zero] at h ⊢ obtain ⟨m, h⟩ := h diff --git a/Mathlib/FieldTheory/Relrank.lean b/Mathlib/FieldTheory/Relrank.lean index 6fe6ff6b4ca8b..4c7805692e150 100644 --- a/Mathlib/FieldTheory/Relrank.lean +++ b/Mathlib/FieldTheory/Relrank.lean @@ -35,11 +35,19 @@ variable {E : Type v} [Field E] {L : Type w} [Field L] variable (A B C : Subfield E) +#adaptation_note +/-- +This `synthInstance.maxHeartbeats` (and below) was required after nightly-2024-11-14; +it's not exactly clear why, but we were very close to the limit previously, +so probably we should not particularly blame changes in Lean, and instead optimize in Mathlib. +-/ +set_option synthInstance.maxHeartbeats 400000 in /-- `Subfield.relrank A B` is defined to be `[B : A ⊓ B]` as a `Cardinal`, in particular, when `A ≤ B` it is `[B : A]`, the degree of the field extension `B / A`. This is similar to `Subgroup.relindex` but it is `Cardinal` valued. -/ noncomputable def relrank := Module.rank ↥(A ⊓ B) (extendScalars (inf_le_right : A ⊓ B ≤ B)) +set_option synthInstance.maxHeartbeats 400000 in /-- The `Nat` version of `Subfield.relrank`. If `B / A ⊓ B` is an infinite extension, then it is zero. -/ noncomputable def relfinrank := finrank ↥(A ⊓ B) (extendScalars (inf_le_right : A ⊓ B ≤ B)) @@ -55,12 +63,14 @@ theorem relrank_eq_of_inf_eq (h : A ⊓ C = B ⊓ C) : relrank A C = relrank B C theorem relfinrank_eq_of_inf_eq (h : A ⊓ C = B ⊓ C) : relfinrank A C = relfinrank B C := congr(toNat $(relrank_eq_of_inf_eq h)) +set_option synthInstance.maxHeartbeats 400000 in /-- If `A ≤ B`, then `Subfield.relrank A B` is `[B : A]` -/ theorem relrank_eq_rank_of_le (h : A ≤ B) : relrank A B = Module.rank A (extendScalars h) := by rw [relrank] have := inf_of_le_left h congr! +set_option synthInstance.maxHeartbeats 400000 in /-- If `A ≤ B`, then `Subfield.relfinrank A B` is `[B : A]` -/ theorem relfinrank_eq_finrank_of_le (h : A ≤ B) : relfinrank A B = finrank A (extendScalars h) := congr(toNat $(relrank_eq_rank_of_le h)) @@ -112,6 +122,7 @@ theorem relrank_top_left : relrank ⊤ A = 1 := relrank_eq_one_of_le le_top @[simp] theorem relfinrank_top_left : relfinrank ⊤ A = 1 := relfinrank_eq_one_of_le le_top +set_option synthInstance.maxHeartbeats 400000 in @[simp] theorem relrank_top_right : relrank A ⊤ = Module.rank A E := by rw [relrank_eq_rank_of_le (show A ≤ ⊤ from le_top), extendScalars_top, diff --git a/Mathlib/FieldTheory/SeparableDegree.lean b/Mathlib/FieldTheory/SeparableDegree.lean index c6a67370781a2..f284a0edd9ea5 100644 --- a/Mathlib/FieldTheory/SeparableDegree.lean +++ b/Mathlib/FieldTheory/SeparableDegree.lean @@ -778,7 +778,7 @@ theorem finSepDegree_eq_finrank_iff [FiniteDimensional F E] : have halg := IsAlgebraic.of_finite F x refine (finSepDegree_adjoin_simple_eq_finrank_iff F E x halg).1 <| le_antisymm (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' + have := Nat.mul_lt_mul_of_lt_of_le' h (finSepDegree_le_finrank F⟮x⟯ E) Fin.pos' rw [finSepDegree_mul_finSepDegree_of_isAlgebraic F F⟮x⟯ E, Module.finrank_mul_finrank F F⟮x⟯ E] at this linarith only [heq, this]⟩, fun _ ↦ finSepDegree_eq_finrank_of_isSeparable F E⟩ diff --git a/Mathlib/Geometry/Euclidean/MongePoint.lean b/Mathlib/Geometry/Euclidean/MongePoint.lean index c3bf888133c63..7a59b661c7066 100644 --- a/Mathlib/Geometry/Euclidean/MongePoint.lean +++ b/Mathlib/Geometry/Euclidean/MongePoint.lean @@ -442,12 +442,8 @@ theorem orthocenter_eq_of_range_eq {t₁ t₂ : Triangle ℝ P} planes. -/ theorem altitude_eq_mongePlane (t : Triangle ℝ P) {i₁ i₂ i₃ : Fin 3} (h₁₂ : i₁ ≠ i₂) (h₁₃ : i₁ ≠ i₃) (h₂₃ : i₂ ≠ i₃) : t.altitude i₁ = t.mongePlane i₂ i₃ := by - have hs : ({i₂, i₃}ᶜ : Finset (Fin 3)) = {i₁} := by - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11043): was `decide!` - revert i₁ i₂ i₃; decide - have he : univ.erase i₁ = {i₂, i₃} := by - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11043): was `decide!` - revert i₁ i₂ i₃; decide + have hs : ({i₂, i₃}ᶜ : Finset (Fin 3)) = {i₁} := by decide +revert + have he : univ.erase i₁ = {i₂, i₃} := by decide +revert rw [mongePlane_def, altitude_def, direction_affineSpan, hs, he, centroid_singleton, coe_insert, coe_singleton, vectorSpan_image_eq_span_vsub_set_left_ne ℝ _ (Set.mem_insert i₂ _)] simp [h₂₃, Submodule.span_insert_eq_span] @@ -456,8 +452,7 @@ theorem altitude_eq_mongePlane (t : Triangle ℝ P) {i₁ i₂ i₃ : Fin 3} (h theorem orthocenter_mem_altitude (t : Triangle ℝ P) {i₁ : Fin 3} : t.orthocenter ∈ t.altitude i₁ := by obtain ⟨i₂, i₃, h₁₂, h₂₃, h₁₃⟩ : ∃ i₂ i₃, i₁ ≠ i₂ ∧ i₂ ≠ i₃ ∧ i₁ ≠ i₃ := by - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11043): was `decide!` - fin_cases i₁ <;> decide + decide +revert rw [orthocenter_eq_mongePoint, t.altitude_eq_mongePlane h₁₂ h₁₃ h₂₃] exact t.mongePoint_mem_mongePlane @@ -467,8 +462,7 @@ theorem eq_orthocenter_of_forall_mem_altitude {t : Triangle ℝ P} {i₁ i₂ : (h₁₂ : i₁ ≠ i₂) (h₁ : p ∈ t.altitude i₁) (h₂ : p ∈ t.altitude i₂) : p = t.orthocenter := by obtain ⟨i₃, h₂₃, h₁₃⟩ : ∃ i₃, i₂ ≠ i₃ ∧ i₁ ≠ i₃ := by clear h₁ h₂ - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11043): was `decide!` - fin_cases i₁ <;> fin_cases i₂ <;> decide + decide +revert rw [t.altitude_eq_mongePlane h₁₃ h₁₂ h₂₃.symm] at h₁ rw [t.altitude_eq_mongePlane h₂₃ h₁₂.symm h₁₃.symm] at h₂ rw [orthocenter_eq_mongePoint] @@ -476,8 +470,7 @@ theorem eq_orthocenter_of_forall_mem_altitude {t : Triangle ℝ P} {i₁ i₂ : intro i hi have hi₁₂ : i₁ = i ∨ i₂ = i := by clear h₁ h₂ - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11043): was `decide!` - fin_cases i₁ <;> fin_cases i₂ <;> fin_cases i₃ <;> fin_cases i <;> simp at h₁₂ h₁₃ h₂₃ hi ⊢ + decide +revert cases' hi₁₂ with hi₁₂ hi₁₂ · exact hi₁₂ ▸ h₂ · exact hi₁₂ ▸ h₁ @@ -498,8 +491,7 @@ theorem dist_orthocenter_reflection_circumcenter (t : Triangle ℝ P) {i₁ i₂ have hu : ({i₁, i₂} : Finset (Fin 3)) ⊆ univ := subset_univ _ obtain ⟨i₃, hi₃, hi₃₁, hi₃₂⟩ : ∃ i₃, univ \ ({i₁, i₂} : Finset (Fin 3)) = {i₃} ∧ i₃ ≠ i₁ ∧ i₃ ≠ i₂ := by - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11043): was `decide!` - fin_cases i₁ <;> fin_cases i₂ <;> simp at h <;> decide + decide +revert simp_rw [← sum_sdiff hu, hi₃] norm_num [hi₃₁, hi₃₂] @@ -542,8 +534,7 @@ theorem altitude_replace_orthocenter_eq_affineSpan {t₁ t₂ : Triangle ℝ P} ?_ · have hu : (Finset.univ : Finset (Fin 3)) = {j₁, j₂, j₃} := by clear h₁ h₂ h₃ - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11043): was `decide!` - revert i₁ i₂ i₃ j₁ j₂ j₃; decide + decide +revert rw [← Set.image_univ, ← Finset.coe_univ, hu, Finset.coe_insert, Finset.coe_insert, Finset.coe_singleton, Set.image_insert_eq, Set.image_insert_eq, Set.image_singleton, h₁, h₂, h₃, Set.insert_subset_iff, Set.insert_subset_iff, Set.singleton_subset_iff] @@ -557,16 +548,14 @@ theorem altitude_replace_orthocenter_eq_affineSpan {t₁ t₂ : Triangle ℝ P} use mem_affineSpan ℝ (Set.mem_range_self _) have hu : Finset.univ.erase j₂ = {j₁, j₃} := by clear h₁ h₂ h₃ - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11043): was `decide!` - revert i₁ i₂ i₃ j₁ j₂ j₃; decide + decide +revert rw [hu, Finset.coe_insert, Finset.coe_singleton, Set.image_insert_eq, Set.image_singleton, h₁, h₃] have hle : (t₁.altitude i₃).directionᗮ ≤ line[ℝ, t₁.orthocenter, t₁.points i₃].directionᗮ := Submodule.orthogonal_le (direction_le (affineSpan_orthocenter_point_le_altitude _ _)) refine hle ((t₁.vectorSpan_isOrtho_altitude_direction i₃) ?_) have hui : Finset.univ.erase i₃ = {i₁, i₂} := by clear hle h₂ h₃ - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11043): was `decide!` - revert i₁ i₂ i₃; decide + decide +revert rw [hui, Finset.coe_insert, Finset.coe_singleton, Set.image_insert_eq, Set.image_singleton] exact vsub_mem_vectorSpan ℝ (Set.mem_insert _ _) (Set.mem_insert_of_mem _ (Set.mem_singleton _)) @@ -622,7 +611,7 @@ theorem exists_of_range_subset_orthocentricSystem {t : Triangle ℝ P} obtain ⟨i₂, i₃, h₁₂, h₁₃, h₂₃, h₁₂₃⟩ : ∃ i₂ i₃ : Fin 3, i₁ ≠ i₂ ∧ i₁ ≠ i₃ ∧ i₂ ≠ i₃ ∧ ∀ i : Fin 3, i = i₁ ∨ i = i₂ ∨ i = i₃ := by clear h₁ - fin_cases i₁ <;> decide + decide +revert have h : ∀ i, i₁ ≠ i → ∃ j : Fin 3, t.points j = p i := by intro i hi replace hps := Set.mem_of_mem_insert_of_ne @@ -729,8 +718,7 @@ theorem OrthocentricSystem.eq_insert_orthocenter {s : Set P} (ho : OrthocentricS · obtain ⟨j₁, hj₁₂, hj₁₃, hj₁₂₃⟩ : ∃ j₁ : Fin 3, j₁ ≠ j₂ ∧ j₁ ≠ j₃ ∧ ∀ j : Fin 3, j = j₁ ∨ j = j₂ ∨ j = j₃ := by clear h₂ h₃ - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11043): was `decide!` - revert j₂ j₃; decide + decide +revert suffices h : t₀.points j₁ = t.orthocenter by have hui : (Set.univ : Set (Fin 3)) = {i₁, i₂, i₃} := by ext x; simpa using h₁₂₃ x have huj : (Set.univ : Set (Fin 3)) = {j₁, j₂, j₃} := by ext x; simpa using hj₁₂₃ x diff --git a/Mathlib/Geometry/Manifold/Instances/Real.lean b/Mathlib/Geometry/Manifold/Instances/Real.lean index ce5fb5a376760..8e667f22b2b24 100644 --- a/Mathlib/Geometry/Manifold/Instances/Real.lean +++ b/Mathlib/Geometry/Manifold/Instances/Real.lean @@ -98,6 +98,7 @@ theorem interior_halfSpace {n : ℕ} (p : ℝ≥0∞) (a : ℝ) (i : Fin n) : change interior (f ⁻¹' Ici a) = f ⁻¹' Ioi a rw [f.interior_preimage, interior_Ici] apply Function.surjective_eval + @[deprecated (since := "2024-11-12")] alias interior_halfspace := interior_halfSpace open ENNReal in @@ -108,6 +109,7 @@ theorem closure_halfSpace {n : ℕ} (p : ℝ≥0∞) (a : ℝ) (i : Fin n) : change closure (f ⁻¹' Ici a) = f ⁻¹' Ici a rw [f.closure_preimage, closure_Ici] apply Function.surjective_eval + @[deprecated (since := "2024-11-12")] alias closure_halfspace := closure_halfSpace open ENNReal in @@ -118,6 +120,7 @@ theorem closure_open_halfSpace {n : ℕ} (p : ℝ≥0∞) (a : ℝ) (i : Fin n) change closure (f ⁻¹' Ioi a) = f ⁻¹' Ici a rw [f.closure_preimage, closure_Ioi] apply Function.surjective_eval + @[deprecated (since := "2024-11-12")] alias closure_open_halfspace := closure_open_halfSpace open ENNReal in diff --git a/Mathlib/Geometry/Manifold/MFDeriv/NormedSpace.lean b/Mathlib/Geometry/Manifold/MFDeriv/NormedSpace.lean index 7b650ffa5716c..b30e07d620aa0 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/NormedSpace.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/NormedSpace.lean @@ -81,19 +81,33 @@ end Module /-! ### Linear maps between normed spaces are smooth -/ +#adaptation_note +/-- +After https://github.com/leanprover/lean4/pull/6024 +we needed to add the named arguments `𝕜 := 𝕜` and `F := ((F₂ →L[𝕜] F₃) →L[𝕜] F₁ →L[𝕜] F₃))` +to `ContinuousLinearMap.differentiable`. +-/ theorem MDifferentiableWithinAt.clm_precomp {f : M → F₁ →L[𝕜] F₂} {s : Set M} {x : M} (hf : MDifferentiableWithinAt I 𝓘(𝕜, F₁ →L[𝕜] F₂) f s x) : MDifferentiableWithinAt I 𝓘(𝕜, (F₂ →L[𝕜] F₃) →L[𝕜] (F₁ →L[𝕜] F₃)) (fun y ↦ (f y).precomp F₃ : M → (F₂ →L[𝕜] F₃) →L[𝕜] (F₁ →L[𝕜] F₃)) s x := Differentiable.comp_mdifferentiableWithinAt (g := (ContinuousLinearMap.compL 𝕜 F₁ F₂ F₃).flip) - (ContinuousLinearMap.compL 𝕜 F₁ F₂ F₃).flip.differentiable hf - + (ContinuousLinearMap.differentiable (𝕜 := 𝕜) (F := ((F₂ →L[𝕜] F₃) →L[𝕜] F₁ →L[𝕜] F₃)) + (ContinuousLinearMap.compL 𝕜 F₁ F₂ F₃).flip) hf + +#adaptation_note +/-- +After https://github.com/leanprover/lean4/pull/6024 +we needed to add the named arguments `𝕜 := 𝕜` and `F := ((F₂ →L[𝕜] F₃) →L[𝕜] F₁ →L[𝕜] F₃))` +to `ContinuousLinearMap.differentiable`. +-/ nonrec theorem MDifferentiableAt.clm_precomp {f : M → F₁ →L[𝕜] F₂} {x : M} (hf : MDifferentiableAt I 𝓘(𝕜, F₁ →L[𝕜] F₂) f x) : MDifferentiableAt I 𝓘(𝕜, (F₂ →L[𝕜] F₃) →L[𝕜] (F₁ →L[𝕜] F₃)) (fun y ↦ (f y).precomp F₃ : M → (F₂ →L[𝕜] F₃) →L[𝕜] (F₁ →L[𝕜] F₃)) x := Differentiable.comp_mdifferentiableAt (g := (ContinuousLinearMap.compL 𝕜 F₁ F₂ F₃).flip) - (ContinuousLinearMap.compL 𝕜 F₁ F₂ F₃).flip.differentiable hf + (ContinuousLinearMap.differentiable (𝕜 := 𝕜) (F := ((F₂ →L[𝕜] F₃) →L[𝕜] F₁ →L[𝕜] F₃)) + (ContinuousLinearMap.compL 𝕜 F₁ F₂ F₃).flip) hf theorem MDifferentiableOn.clm_precomp {f : M → F₁ →L[𝕜] F₂} {s : Set M} (hf : MDifferentiableOn I 𝓘(𝕜, F₁ →L[𝕜] F₂) f s) : @@ -107,21 +121,35 @@ theorem MDifferentiable.clm_precomp (fun y ↦ (f y).precomp F₃ : M → (F₂ →L[𝕜] F₃) →L[𝕜] (F₁ →L[𝕜] F₃)) := fun x ↦ (hf x).clm_precomp +#adaptation_note +/-- +After https://github.com/leanprover/lean4/pull/6024 +we needed to add the named arguments `𝕜 := 𝕜` and `F := ((F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₃))` +to `ContinuousLinearMap.differentiable`. +-/ theorem MDifferentiableWithinAt.clm_postcomp {f : M → F₂ →L[𝕜] F₃} {s : Set M} {x : M} (hf : MDifferentiableWithinAt I 𝓘(𝕜, F₂ →L[𝕜] F₃) f s x) : MDifferentiableWithinAt I 𝓘(𝕜, (F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₃)) (fun y ↦ (f y).postcomp F₁ : M → (F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₃)) s x := Differentiable.comp_mdifferentiableWithinAt (F' := (F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₃)) (g := ContinuousLinearMap.compL 𝕜 F₁ F₂ F₃) - (ContinuousLinearMap.compL 𝕜 F₁ F₂ F₃).differentiable hf - + (ContinuousLinearMap.differentiable (𝕜 := 𝕜) (F := ((F₁ →L[𝕜] F₂) →L[𝕜] F₁ →L[𝕜] F₃)) + (ContinuousLinearMap.compL 𝕜 F₁ F₂ F₃)) hf + +#adaptation_note +/-- +After https://github.com/leanprover/lean4/pull/6024 +we needed to add the named arguments `𝕜 := 𝕜` and `F := ((F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₃))` +to `ContinuousLinearMap.differentiable`. +-/ theorem MDifferentiableAt.clm_postcomp {f : M → F₂ →L[𝕜] F₃} {x : M} (hf : MDifferentiableAt I 𝓘(𝕜, F₂ →L[𝕜] F₃) f x) : MDifferentiableAt I 𝓘(𝕜, (F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₃)) (fun y ↦ (f y).postcomp F₁ : M → (F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₃)) x := Differentiable.comp_mdifferentiableAt (F' := (F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₃)) (g := ContinuousLinearMap.compL 𝕜 F₁ F₂ F₃) - (ContinuousLinearMap.compL 𝕜 F₁ F₂ F₃).differentiable hf + (ContinuousLinearMap.differentiable (𝕜 := 𝕜) (F := ((F₁ →L[𝕜] F₂) →L[𝕜] F₁ →L[𝕜] F₃)) + (ContinuousLinearMap.compL 𝕜 F₁ F₂ F₃)) hf nonrec theorem MDifferentiableOn.clm_postcomp {f : M → F₂ →L[𝕜] F₃} {s : Set M} (hf : MDifferentiableOn I 𝓘(𝕜, F₂ →L[𝕜] F₃) f s) : diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean index 5c87be9ceb0bc..df614fe6a7795 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean @@ -92,7 +92,7 @@ that the `U i`'s are open subspaces of the glued space. -/ -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. -- @[nolint has_nonempty_instance] -structure GlueData extends GlueData (PresheafedSpace.{u, v, v} C) where +structure GlueData extends CategoryTheory.GlueData (PresheafedSpace.{u, v, v} C) where f_open : ∀ i j, IsOpenImmersion (f i j) attribute [instance] GlueData.f_open diff --git a/Mathlib/GroupTheory/Congruence/Basic.lean b/Mathlib/GroupTheory/Congruence/Basic.lean index 259eef9b6e583..5f5f5d7432161 100644 --- a/Mathlib/GroupTheory/Congruence/Basic.lean +++ b/Mathlib/GroupTheory/Congruence/Basic.lean @@ -256,7 +256,7 @@ lemma comapQuotientEquivOfSurj_symm_mk (c : Con M) {f : N →* M} (hf) (x : N) : MulEquiv function"] lemma comapQuotientEquivOfSurj_symm_mk' (c : Con M) (f : N ≃* M) (x : N) : ((@MulEquiv.symm (Con.Quotient (comap ⇑f _ c)) _ _ _ - (comapQuotientEquivOfSurj c f f.surjective)) ⟦f x⟧) = ↑x := + (comapQuotientEquivOfSurj c (f : N →* M) f.surjective)) ⟦f x⟧) = ↑x := (MulEquiv.symm_apply_eq (@comapQuotientEquivOfSurj M N _ _ c f _)).mpr rfl /-- The **second isomorphism theorem for monoids**. -/ diff --git a/Mathlib/GroupTheory/Coxeter/Inversion.lean b/Mathlib/GroupTheory/Coxeter/Inversion.lean index c5c613cf20c6a..6afee6ebfd411 100644 --- a/Mathlib/GroupTheory/Coxeter/Inversion.lean +++ b/Mathlib/GroupTheory/Coxeter/Inversion.lean @@ -277,7 +277,7 @@ theorem getD_rightInvSeq_mul_self (ω : List B) (j : ℕ) : rcases em (j < ω.length) with hj | nhj · rw [get?_eq_get hj] simp [← mul_assoc] - · rw [get?_eq_none.mpr (by omega)] + · rw [get?_eq_none_iff.mpr (by omega)] simp theorem getD_leftInvSeq_mul_self (ω : List B) (j : ℕ) : @@ -286,7 +286,7 @@ theorem getD_leftInvSeq_mul_self (ω : List B) (j : ℕ) : rcases em (j < ω.length) with hj | nhj · rw [get?_eq_get hj] simp [← mul_assoc] - · rw [get?_eq_none.mpr (by omega)] + · rw [get?_eq_none_iff.mpr (by omega)] simp theorem rightInvSeq_drop (ω : List B) (j : ℕ) : diff --git a/Mathlib/GroupTheory/Perm/Cycle/Type.lean b/Mathlib/GroupTheory/Perm/Cycle/Type.lean index 7b8e2f1c0b26d..b47673b955137 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Type.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Type.lean @@ -394,12 +394,12 @@ section Cauchy variable (G : Type*) [Group G] (n : ℕ) /-- The type of vectors with terms from `G`, length `n`, and product equal to `1:G`. -/ -def vectorsProdEqOne : Set (Vector G n) := +def vectorsProdEqOne : Set (Mathlib.Vector G n) := { v | v.toList.prod = 1 } namespace VectorsProdEqOne -theorem mem_iff {n : ℕ} (v : Vector G n) : v ∈ vectorsProdEqOne G n ↔ v.toList.prod = 1 := +theorem mem_iff {n : ℕ} (v : Mathlib.Vector G n) : v ∈ vectorsProdEqOne G n ↔ v.toList.prod = 1 := Iff.rfl theorem zero_eq : vectorsProdEqOne G 0 = {Vector.nil} := @@ -421,7 +421,7 @@ instance oneUnique : Unique (vectorsProdEqOne G 1) := by /-- Given a vector `v` of length `n`, make a vector of length `n + 1` whose product is `1`, by appending the inverse of the product of `v`. -/ @[simps] -def vectorEquiv : Vector G n ≃ vectorsProdEqOne G (n + 1) where +def vectorEquiv : Mathlib.Vector G n ≃ vectorsProdEqOne G (n + 1) where toFun v := ⟨v.toList.prod⁻¹ ::ᵥ v, by rw [mem_iff, Vector.toList_cons, List.prod_cons, inv_mul_cancel]⟩ invFun v := v.1.tail @@ -436,12 +436,12 @@ def vectorEquiv : Vector G n ≃ vectorsProdEqOne G (n + 1) where /-- Given a vector `v` of length `n` whose product is 1, make a vector of length `n - 1`, by deleting the last entry of `v`. -/ -def equivVector : ∀ n, vectorsProdEqOne G n ≃ Vector G (n - 1) +def equivVector : ∀ n, vectorsProdEqOne G n ≃ Mathlib.Vector G (n - 1) | 0 => (equivOfUnique (vectorsProdEqOne G 0) (vectorsProdEqOne G 1)).trans (vectorEquiv G 0).symm | (n + 1) => (vectorEquiv G n).symm instance [Fintype G] : Fintype (vectorsProdEqOne G n) := - Fintype.ofEquiv (Vector G (n - 1)) (equivVector G n).symm + Fintype.ofEquiv (Mathlib.Vector G (n - 1)) (equivVector G n).symm theorem card [Fintype G] : Fintype.card (vectorsProdEqOne G n) = Fintype.card G ^ (n - 1) := (Fintype.card_congr (equivVector G n)).trans (card_vector (n - 1)) diff --git a/Mathlib/Lean/Meta/DiscrTree.lean b/Mathlib/Lean/Meta/DiscrTree.lean index 4698847cefc5a..586b3f28ad8c3 100644 --- a/Mathlib/Lean/Meta/DiscrTree.lean +++ b/Mathlib/Lean/Meta/DiscrTree.lean @@ -24,21 +24,21 @@ so that we return lemmas matching larger subexpressions first, and amongst those we return more specific lemmas first. -/ partial def getSubexpressionMatches {α : Type} - (d : DiscrTree α) (e : Expr) (config : WhnfCoreConfig) : MetaM (Array α) := do + (d : DiscrTree α) (e : Expr) : MetaM (Array α) := do match e with | .bvar _ => return #[] | .forallE _ _ _ _ => forallTelescope e (fun args body => do args.foldlM (fun acc arg => do - pure <| acc ++ (← d.getSubexpressionMatches (← inferType arg) config)) - (← d.getSubexpressionMatches body config).reverse) + pure <| acc ++ (← d.getSubexpressionMatches (← inferType arg))) + (← d.getSubexpressionMatches body).reverse) | .lam _ _ _ _ | .letE _ _ _ _ _ => lambdaLetTelescope e (fun args body => do args.foldlM (fun acc arg => do - pure <| acc ++ (← d.getSubexpressionMatches (← inferType arg) config)) - (← d.getSubexpressionMatches body config).reverse) + pure <| acc ++ (← d.getSubexpressionMatches (← inferType arg))) + (← d.getSubexpressionMatches body).reverse) | _ => e.foldlM (fun a f => do - pure <| a ++ (← d.getSubexpressionMatches f config)) (← d.getMatch e config).reverse + pure <| a ++ (← d.getSubexpressionMatches f)) (← d.getMatch e).reverse /-- Check if a `keys : Array DiscTree.Key` is "specific", diff --git a/Mathlib/Lean/Meta/KAbstractPositions.lean b/Mathlib/Lean/Meta/KAbstractPositions.lean index 60d3ded29a824..ffb5899f96568 100644 --- a/Mathlib/Lean/Meta/KAbstractPositions.lean +++ b/Mathlib/Lean/Meta/KAbstractPositions.lean @@ -71,7 +71,7 @@ def viewKAbstractSubExpr (e : Expr) (pos : SubExpr.Pos) : MetaM (Option (Expr × if subExpr.hasLooseBVars then return none let positions ← kabstractPositions subExpr e - let some n := positions.getIdx? pos | unreachable! + let some n := positions.indexOf? pos | unreachable! return some (subExpr, if positions.size == 1 then none else some (n + 1)) /-- Determine whether the result of abstracting `subExpr` from `e` at position `pos` results diff --git a/Mathlib/Lean/Meta/Simp.lean b/Mathlib/Lean/Meta/Simp.lean index 5f38f103a0dcc..289c25c8f0e60 100644 --- a/Mathlib/Lean/Meta/Simp.lean +++ b/Mathlib/Lean/Meta/Simp.lean @@ -72,10 +72,10 @@ def simpTheoremsOfNames (lemmas : List Name := []) (simpOnly : Bool := false) : /-- Construct a `Simp.Context` from a list of names. -/ def Simp.Context.ofNames (lemmas : List Name := []) (simpOnly : Bool := false) - (config : Simp.Config := {}) : MetaM Simp.Context := do pure <| - { simpTheorems := #[← simpTheoremsOfNames lemmas simpOnly], - congrTheorems := ← Lean.Meta.getSimpCongrTheorems, - config := config } + (config : Simp.Config := {}) : MetaM Simp.Context := do + Simp.mkContext config + (simpTheorems := #[← simpTheoremsOfNames lemmas simpOnly]) + (congrTheorems := ← Lean.Meta.getSimpCongrTheorems) /-- Simplify an expression using only a list of lemmas specified by name. -/ def simpOnlyNames (lemmas : List Name) (e : Expr) (config : Simp.Config := {}) : diff --git a/Mathlib/LinearAlgebra/Prod.lean b/Mathlib/LinearAlgebra/Prod.lean index 23732c0a35284..33f885c1bc407 100644 --- a/Mathlib/LinearAlgebra/Prod.lean +++ b/Mathlib/LinearAlgebra/Prod.lean @@ -985,7 +985,8 @@ most once. Then the image of `f` is the graph of some linear map `f' : H → I`. lemma LinearMap.exists_range_eq_graph {f : G →ₛₗ[σ] H × I} (hf₁ : Surjective (Prod.fst ∘ f)) (hf : ∀ g₁ g₂, (f g₁).1 = (f g₂).1 → (f g₁).2 = (f g₂).2) : ∃ f' : H →ₗ[S] I, LinearMap.range f = LinearMap.graph f' := by - obtain ⟨f', hf'⟩ := AddMonoidHom.exists_mrange_eq_mgraph (I := I) (f := f) hf₁ hf + obtain ⟨f', hf'⟩ := + AddMonoidHom.exists_mrange_eq_mgraph (G := G) (H := H) (I := I) (f := f) hf₁ hf simp only [SetLike.ext_iff, AddMonoidHom.mem_mrange, AddMonoidHom.coe_coe, AddMonoidHom.mem_mgraph] at hf' use diff --git a/Mathlib/LinearAlgebra/TensorProduct/Basic.lean b/Mathlib/LinearAlgebra/TensorProduct/Basic.lean index 231c9e14ba44f..0a966b059726d 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Basic.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Basic.lean @@ -91,6 +91,9 @@ namespace TensorProduct section Module +protected instance zero : Zero (M ⊗[R] N) := + (addConGen (TensorProduct.Eqv R M N)).zero + protected instance add : Add (M ⊗[R] N) := (addConGen (TensorProduct.Eqv R M N)).hasAdd @@ -99,7 +102,8 @@ instance addZeroClass : AddZeroClass (M ⊗[R] N) := /- The `toAdd` field is given explicitly as `TensorProduct.add` for performance reasons. This avoids any need to unfold `Con.addMonoid` when the type checker is checking that instance diagrams commute -/ - toAdd := TensorProduct.add _ _ } + toAdd := TensorProduct.add _ _ + toZero := TensorProduct.zero _ _ } instance addSemigroup : AddSemigroup (M ⊗[R] N) := { (addConGen (TensorProduct.Eqv R M N)).addMonoid with @@ -320,7 +324,7 @@ protected theorem add_smul (r s : R'') (x : M ⊗[R] N) : (r + s) • x = r • instance addMonoid : AddMonoid (M ⊗[R] N) := { TensorProduct.addZeroClass _ _ with toAddSemigroup := TensorProduct.addSemigroup _ _ - toZero := (TensorProduct.addZeroClass _ _).toZero + toZero := TensorProduct.zero _ _ nsmul := fun n v => n • v nsmul_zero := by simp [TensorProduct.zero_smul] nsmul_succ := by simp only [TensorProduct.one_smul, TensorProduct.add_smul, add_comm, diff --git a/Mathlib/Logic/Equiv/List.lean b/Mathlib/Logic/Equiv/List.lean index c0753898bf0dc..6a6bb07e4ed21 100644 --- a/Mathlib/Logic/Equiv/List.lean +++ b/Mathlib/Logic/Equiv/List.lean @@ -127,11 +127,11 @@ noncomputable def _root_.Fintype.toEncodable (α : Type*) [Fintype α] : Encodab classical exact (Fintype.truncEncodable α).out /-- If `α` is encodable, then so is `Vector α n`. -/ -instance _root_.Vector.encodable [Encodable α] {n} : Encodable (Vector α n) := +instance Mathlib.Vector.encodable [Encodable α] {n} : Encodable (Mathlib.Vector α n) := Subtype.encodable /-- If `α` is countable, then so is `Vector α n`. -/ -instance _root_.Vector.countable [Countable α] {n} : Countable (Vector α n) := +instance Mathlib.Vector.countable [Countable α] {n} : Countable (Mathlib.Vector α n) := Subtype.countable /-- If `α` is encodable, then so is `Fin n → α`. -/ diff --git a/Mathlib/Logic/Small/List.lean b/Mathlib/Logic/Small/List.lean index 13b6fd0b0469a..f0706af610d10 100644 --- a/Mathlib/Logic/Small/List.lean +++ b/Mathlib/Logic/Small/List.lean @@ -18,9 +18,9 @@ universe u v open Mathlib -instance smallVector {α : Type v} {n : ℕ} [Small.{u} α] : Small.{u} (Vector α n) := +instance smallVector {α : Type v} {n : ℕ} [Small.{u} α] : Small.{u} (Mathlib.Vector α n) := small_of_injective (Equiv.vectorEquivFin α n).injective instance smallList {α : Type v} [Small.{u} α] : Small.{u} (List α) := by - let e : (Σn, Vector α n) ≃ List α := Equiv.sigmaFiberEquiv List.length + let e : (Σn, Mathlib.Vector α n) ≃ List α := Equiv.sigmaFiberEquiv List.length exact small_of_surjective e.surjective diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index ca21a90b697cf..1d8f510752db3 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -1508,9 +1508,6 @@ private theorem lintegral_rpow_sum_coe_nnnorm_sub_le_rpow_tsum · rw [Real.norm_of_nonneg _] exact Finset.sum_nonneg fun x _ => norm_nonneg _ · exact fun x _ => norm_nonneg _ - change - (∫⁻ a, (fun x => ↑‖∑ i ∈ Finset.range (n + 1), ‖f (i + 1) x - f i x‖‖₊ ^ p) a ∂μ) ^ p⁻¹ ≤ - ∑' i, B i at hn rwa [h_nnnorm_nonneg] at hn private theorem lintegral_rpow_tsum_coe_nnnorm_sub_le_tsum {f : ℕ → α → E} diff --git a/Mathlib/MeasureTheory/Group/Measure.lean b/Mathlib/MeasureTheory/Group/Measure.lean index 8662ff3b7f53c..449597022125a 100644 --- a/Mathlib/MeasureTheory/Group/Measure.lean +++ b/Mathlib/MeasureTheory/Group/Measure.lean @@ -764,7 +764,12 @@ nonrec theorem _root_.MulEquiv.isHaarMeasure_map [BorelSpace G] [TopologicalGrou [TopologicalGroup H] (e : G ≃* H) (he : Continuous e) (hesymm : Continuous e.symm) : IsHaarMeasure (Measure.map e μ) := let f : G ≃ₜ H := .mk e - isHaarMeasure_map μ e he e.surjective f.isClosedEmbedding.tendsto_cocompact + #adaptation_note + /-- + After https://github.com/leanprover/lean4/pull/6024 + we needed to write `e.toMonoidHom` instead of just `e`, to avoid unification issues. + -/ + isHaarMeasure_map μ e.toMonoidHom he e.surjective f.isClosedEmbedding.tendsto_cocompact /-- A convenience wrapper for MeasureTheory.Measure.isAddHaarMeasure_map`. -/ instance _root_.ContinuousLinearEquiv.isAddHaarMeasure_map diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean b/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean index 62141d5c9d30c..0a63d44706834 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean @@ -506,9 +506,16 @@ def arrowProdEquivProdArrow (α β γ : Type*) [MeasurableSpace α] [MeasurableS __ := Equiv.arrowProdEquivProdArrow α β γ measurable_toFun _ h := by simp_rw [Equiv.arrowProdEquivProdArrow, coe_fn_mk] + #adaptation_note + /-- + After https://github.com/leanprover/lean4/pull/6024 + we need provide the type hints `(a : γ → α × β)`, to avoid unification issues. + -/ exact MeasurableSet.preimage h (Measurable.prod_mk - (measurable_pi_lambda (fun a c ↦ (a c).1) fun a ↦ (measurable_pi_apply a).fst) - (measurable_pi_lambda (fun a c ↦ (a c).2) fun a ↦ (measurable_pi_apply a).snd )) + (measurable_pi_lambda (fun (a : γ → α × β) c ↦ (a c).1) + fun a ↦ (measurable_pi_apply a).fst) + (measurable_pi_lambda (fun (a : γ → α × β) c ↦ (a c).2) + fun a ↦ (measurable_pi_apply a).snd)) measurable_invFun _ h := by simp_rw [Equiv.arrowProdEquivProdArrow, coe_fn_symm_mk] exact MeasurableSet.preimage h (by measurability) diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean index ad3c3db1e4080..fc756040cc012 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean @@ -207,7 +207,7 @@ theorem MeasureTheory.volume_sum_rpow_lt [Nonempty ι] {p : ℝ} (hp : 1 ≤ p) ext x refine ⟨fun hx => ?_, fun hx => hx.elim⟩ exact not_le.mpr (lt_of_lt_of_le (Set.mem_setOf.mp hx) hr) (h₂ x) - rw [this, measure_empty, ← zero_eq_ofReal.mpr hr, zero_pow Fin.size_pos'.ne', zero_mul] + rw [this, measure_empty, ← zero_eq_ofReal.mpr hr, zero_pow Fin.pos'.ne', zero_mul] · rw [← volume_sum_rpow_lt_one _ hp, ← ofReal_pow (le_of_lt hr), ← finrank_pi ℝ] 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, @@ -280,7 +280,7 @@ theorem Complex.volume_sum_rpow_lt [Nonempty ι] {p : ℝ} (hp : 1 ≤ p) (r : ext x refine ⟨fun hx => ?_, fun hx => hx.elim⟩ exact not_le.mpr (lt_of_lt_of_le (Set.mem_setOf.mp hx) hr) (h₂ x) - rw [this, measure_empty, ← zero_eq_ofReal.mpr hr, zero_pow Fin.size_pos'.ne', zero_mul] + rw [this, measure_empty, ← zero_eq_ofReal.mpr hr, zero_pow Fin.pos'.ne', zero_mul] · rw [← Complex.volume_sum_rpow_lt_one _ hp, ← ENNReal.ofReal_pow (le_of_lt hr)] 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, diff --git a/Mathlib/ModelTheory/Encoding.lean b/Mathlib/ModelTheory/Encoding.lean index e3832e469136b..859366d1da80f 100644 --- a/Mathlib/ModelTheory/Encoding.lean +++ b/Mathlib/ModelTheory/Encoding.lean @@ -234,7 +234,7 @@ theorem listDecode_encode_list (l : List (Σn, L.BoundedFormula α n)) : Σn, L.Term (α ⊕ (Fin n)))) (finRange φ_l) ++ l)).get? ↑i).join = some ⟨_, ts i⟩ := by intro i 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] + get?_eq_some_iff, length_append, length_map, length_finRange] refine ⟨lt_of_lt_of_le i.2 le_self_add, ?_⟩ rw [get_eq_getElem, getElem_append_left, getElem_map] · simp only [getElem_finRange, cast_mk, Fin.eta, Function.comp_apply, Sum.getLeft?_inl] diff --git a/Mathlib/ModelTheory/Syntax.lean b/Mathlib/ModelTheory/Syntax.lean index 14e6d78b779dd..ff03d55615d75 100644 --- a/Mathlib/ModelTheory/Syntax.lean +++ b/Mathlib/ModelTheory/Syntax.lean @@ -6,6 +6,7 @@ Authors: Aaron Anderson, Jesse Michael Han, Floris van Doorn import Mathlib.Data.Set.Prod import Mathlib.Logic.Equiv.Fin import Mathlib.ModelTheory.LanguageMap +import Mathlib.Algebra.Order.Ring.Nat /-! # Basics on First-Order Syntax diff --git a/Mathlib/NumberTheory/ClassNumber/AdmissibleAbsoluteValue.lean b/Mathlib/NumberTheory/ClassNumber/AdmissibleAbsoluteValue.lean index bdc2a13623ee7..0d34bda52b147 100644 --- a/Mathlib/NumberTheory/ClassNumber/AdmissibleAbsoluteValue.lean +++ b/Mathlib/NumberTheory/ClassNumber/AdmissibleAbsoluteValue.lean @@ -102,7 +102,7 @@ theorem exists_approx_aux (n : ℕ) (h : abv.IsAdmissible) : /-- This proof was nicer prior to https://github.com/leanprover/lean4/pull/4400. Please feel welcome to improve it, by avoiding use of `List.get` in favour of `GetElem`. -/ have : ∀ i h, t ((Finset.univ.filter fun x ↦ t x = s).toList.get ⟨i, h⟩) = s := fun i h ↦ - (Finset.mem_filter.mp (Finset.mem_toList.mp (List.get_mem _ i h))).2 + (Finset.mem_filter.mp (Finset.mem_toList.mp (List.get_mem _ ⟨i, h⟩))).2 simp only [Nat.succ_eq_add_one, Finset.length_toList, List.get_eq_getElem] at this simp only [Nat.succ_eq_add_one, List.get_eq_getElem, Fin.coe_castLE] rw [this _ (Nat.lt_of_le_of_lt (Nat.le_of_lt_succ i₁.2) hs), diff --git a/Mathlib/NumberTheory/FactorisationProperties.lean b/Mathlib/NumberTheory/FactorisationProperties.lean index 28cabcc7fd958..f32fbb70c0910 100644 --- a/Mathlib/NumberTheory/FactorisationProperties.lean +++ b/Mathlib/NumberTheory/FactorisationProperties.lean @@ -185,7 +185,7 @@ theorem infinite_even_deficient : {n : ℕ | Even n ∧ n.Deficient}.Infinite := constructor · exact ⟨⟨2 ^ n, by ring⟩, prime_two.deficient_pow⟩ · calc - n ≤ 2 ^ n := Nat.le_of_lt (lt_two_pow n) + n ≤ 2 ^ n := Nat.le_of_lt n.lt_two_pow_self _ < 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 diff --git a/Mathlib/NumberTheory/FermatPsp.lean b/Mathlib/NumberTheory/FermatPsp.lean index 31cdde409beb6..a620fe11fa855 100644 --- a/Mathlib/NumberTheory/FermatPsp.lean +++ b/Mathlib/NumberTheory/FermatPsp.lean @@ -314,7 +314,7 @@ private theorem psp_from_prime_gt_p {b : ℕ} (b_ge_two : 2 ≤ b) {p : ℕ} (p_ have : 2 ≤ 2 * p - 2 := le_tsub_of_add_le_left (show 4 ≤ 2 * p by omega) have : 2 + p ≤ 2 * p := by omega have : p ≤ 2 * p - 2 := le_tsub_of_add_le_left this - exact this.trans_lt (lt_pow_self b_ge_two _) + exact this.trans_lt (Nat.lt_pow_self b_ge_two) /-- For all positive bases, there exist infinite **Fermat pseudoprimes** to that base. Given in this form: for all numbers `b ≥ 1` and `m`, there exists a pseudoprime `n` to base `b` such diff --git a/Mathlib/NumberTheory/LSeries/Nonvanishing.lean b/Mathlib/NumberTheory/LSeries/Nonvanishing.lean index 7ed3e3b3b2c37..bcec46ab9e9d3 100644 --- a/Mathlib/NumberTheory/LSeries/Nonvanishing.lean +++ b/Mathlib/NumberTheory/LSeries/Nonvanishing.lean @@ -374,8 +374,13 @@ private lemma LFunction_ne_zero_of_not_quadratic_or_ne_one {t : ℝ} (h : χ ^ 2 -- go via absolute value to translate into a statement over `ℝ` replace H := (H₀.trans H).norm_right simp only [norm_eq_abs, abs_ofReal] at H - exact isLittleO_irrefl (.of_forall (fun _ ↦ one_ne_zero)) <| H.of_norm_right.trans_isLittleO - <| isLittleO_id_one.mono nhdsWithin_le_nhds + #adaptation_note + /-- + After https://github.com/leanprover/lean4/pull/6024 + we needed to add `(F' := ℝ)` to `H.of_norm_right`. + -/ + exact isLittleO_irrefl (.of_forall (fun _ ↦ one_ne_zero)) <| + (H.of_norm_right (F' := ℝ)).trans_isLittleO <| isLittleO_id_one.mono nhdsWithin_le_nhds /-- If `χ` is a Dirichlet character, then `L(χ, s)` does not vanish when `s.re = 1` except when `χ` is trivial and `s = 1` (then `L(χ, s)` has a simple pole at `s = 1`). -/ diff --git a/Mathlib/NumberTheory/Padics/PadicIntegers.lean b/Mathlib/NumberTheory/Padics/PadicIntegers.lean index f9e0e0463f74b..db88a79d9130f 100644 --- a/Mathlib/NumberTheory/Padics/PadicIntegers.lean +++ b/Mathlib/NumberTheory/Padics/PadicIntegers.lean @@ -305,7 +305,7 @@ theorem exists_pow_neg_lt {ε : ℝ} (hε : 0 < ε) : ∃ k : ℕ, (p : ℝ) ^ ( apply lt_of_lt_of_le hk norm_cast apply le_of_lt - convert Nat.lt_pow_self _ _ using 1 + convert Nat.lt_pow_self _ using 1 exact hp.1.one_lt · exact mod_cast hp.1.pos diff --git a/Mathlib/NumberTheory/SmoothNumbers.lean b/Mathlib/NumberTheory/SmoothNumbers.lean index 5072132e5f160..dfc8479b15721 100644 --- a/Mathlib/NumberTheory/SmoothNumbers.lean +++ b/Mathlib/NumberTheory/SmoothNumbers.lean @@ -234,7 +234,7 @@ def equivProdNatFactoredNumbers {s : Finset ℕ} {p : ℕ} (hp : p.Prime) (hs : (filter _ <| perm_primeFactorsList_mul (pow_ne_zero e hp.ne_zero) hm₀).trans ?_ rw [filter_append, hp.primeFactorsList_pow, 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]] + nil_append, filter_eq_self.mpr fun q hq ↦ by simp only [hm q hq, decide_true]] right_inv := by rintro ⟨m, hm₀, hm⟩ simp only [Set.coe_setOf, Set.mem_setOf_eq, Subtype.mk.injEq] @@ -244,8 +244,8 @@ def equivProdNatFactoredNumbers {s : Finset ℕ} {p : ℕ} (hp : p.Prime) (hs : refine (filter_congr fun q hq ↦ ?_).symm simp only [decide_not, Bool.not_eq_true', decide_eq_false_iff_not, decide_eq_true_eq] rcases Finset.mem_insert.mp <| hm _ hq with h | h - · simp only [h, hs, decide_False, Bool.not_false, decide_True] - · simp only [h, decide_True, Bool.not_true, false_eq_decide_iff] + · simp only [h, hs, decide_false, Bool.not_false, decide_true] + · simp only [h, decide_true, Bool.not_true, false_eq_decide_iff] exact fun H ↦ hs <| H ▸ h refine prod_eq <| (filter_eq m.primeFactorsList p).symm ▸ this ▸ perm_append_comm.trans ?_ simp only [decide_not] diff --git a/Mathlib/Order/Filter/Defs.lean b/Mathlib/Order/Filter/Defs.lean index 4cf86091425d0..432915dcbac1d 100644 --- a/Mathlib/Order/Filter/Defs.lean +++ b/Mathlib/Order/Filter/Defs.lean @@ -313,7 +313,7 @@ instance instSProd : SProd (Filter α) (Filter β) (Filter (α × β)) where /-- Product of filters. This is the filter generated by cartesian products of elements of the component filters. Deprecated. Use `f ×ˢ g` instead. -/ -@[deprecated (since := "2024-11-29")] +@[deprecated " Use `f ×ˢ g` instead." (since := "2024-11-29")] protected def prod (f : Filter α) (g : Filter β) : Filter (α × β) := f ×ˢ g theorem prod_eq_inf (f : Filter α) (g : Filter β) : f ×ˢ g = f.comap Prod.fst ⊓ g.comap Prod.snd := diff --git a/Mathlib/Order/Interval/Finset/Box.lean b/Mathlib/Order/Interval/Finset/Box.lean index 1d5823025b14e..9956d1f74bcb2 100644 --- a/Mathlib/Order/Interval/Finset/Box.lean +++ b/Mathlib/Order/Interval/Finset/Box.lean @@ -93,11 +93,6 @@ lemma card_box : ∀ {n}, n ≠ 0 → #(box n : Finset (ℤ × ℤ)) = 8 * n | 0 => by simp [Prod.ext_iff] | n + 1 => by simp [box_succ_eq_sdiff, Prod.le_def] - #adaptation_note /-- v4.7.0-rc1: `omega` no longer identifies atoms up to defeq. - (This had become a performance bottleneck.) - We need a tactic for normalising instances, to avoid the `have`/`simp` dance below: -/ - have : @Nat.cast ℤ instNatCastInt n = @Nat.cast ℤ AddMonoidWithOne.toNatCast n := rfl - simp only [this] omega -- TODO: Can this be generalised to locally finite archimedean ordered rings? diff --git a/Mathlib/Order/WellFounded.lean b/Mathlib/Order/WellFounded.lean index 99a592719d4bc..0511b5790ae02 100644 --- a/Mathlib/Order/WellFounded.lean +++ b/Mathlib/Order/WellFounded.lean @@ -94,21 +94,19 @@ set_option linter.deprecated false open Classical in /-- A successor of an element `x` in a well-founded order is a minimal element `y` such that -`x < y` if one exists. Otherwise it is `x` itself. - -Deprecated. If you have a linear order, consider defining a `SuccOrder` instance through -`ConditionallyCompleteLinearOrder.toSuccOrder`. -/ -@[deprecated (since := "2024-10-25")] +`x < y` if one exists. Otherwise it is `x` itself. -/ +@[deprecated "If you have a linear order, consider defining a `SuccOrder` instance through +`ConditionallyCompleteLinearOrder.toSuccOrder`." (since := "2024-10-25")] protected noncomputable def succ {r : α → α → Prop} (wf : WellFounded r) (x : α) : α := if h : ∃ y, r x y then wf.min { y | r x y } h else x -@[deprecated (since := "2024-10-25")] +@[deprecated "No deprecation message was provided." (since := "2024-10-25")] protected theorem lt_succ {r : α → α → Prop} (wf : WellFounded r) {x : α} (h : ∃ y, r x y) : r x (wf.succ x) := by rw [WellFounded.succ, dif_pos h] apply min_mem -@[deprecated (since := "2024-10-25")] +@[deprecated "No deprecation message was provided." (since := "2024-10-25")] protected theorem lt_succ_iff {r : α → α → Prop} [wo : IsWellOrder α r] {x : α} (h : ∃ y, r x y) (y : α) : r y (wo.wf.succ x) ↔ r y x ∨ y = x := by constructor diff --git a/Mathlib/RingTheory/Algebraic/Basic.lean b/Mathlib/RingTheory/Algebraic/Basic.lean index 98fbcdfdd5f93..2768c2dca2d4d 100644 --- a/Mathlib/RingTheory/Algebraic/Basic.lean +++ b/Mathlib/RingTheory/Algebraic/Basic.lean @@ -195,7 +195,7 @@ theorem IsAlgebraic.of_ringHom_of_comp_eq (halg : IsAlgebraic S (g a)) (h : RingHom.comp (algebraMap S B) f = RingHom.comp g (algebraMap R A)) : IsAlgebraic R a := by obtain ⟨p, h1, h2⟩ := halg - obtain ⟨q, rfl⟩ := map_surjective f hf p + obtain ⟨q, rfl⟩ := map_surjective (f : R →+* S) hf p refine ⟨q, fun h' ↦ by simp [h'] at h1, hg ?_⟩ change aeval ((g : A →+* B) a) _ = 0 at h2 change (g : A →+* B) _ = _ diff --git a/Mathlib/RingTheory/AlgebraicIndependent.lean b/Mathlib/RingTheory/AlgebraicIndependent.lean index ba9cd9de0655c..21ed4dbc29d3f 100644 --- a/Mathlib/RingTheory/AlgebraicIndependent.lean +++ b/Mathlib/RingTheory/AlgebraicIndependent.lean @@ -276,7 +276,7 @@ theorem AlgebraicIndependent.of_ringHom_of_comp_eq (H : AlgebraicIndependent S ( have := H (p.map f) <| by have : (g : A →+* B) _ = _ := congr(g $hp) rwa [map_zero, map_aeval, ← h, ← eval₂Hom_map_hom, ← aeval_eq_eval₂Hom] at this - exact map_injective f hf (by rwa [map_zero]) + exact map_injective (f : R →+* S) hf (by rwa [map_zero]) theorem AlgebraicIndependent.ringHom_of_comp_eq (H : AlgebraicIndependent R x) (hf : Function.Surjective f) (hg : Function.Injective g) @@ -284,7 +284,7 @@ theorem AlgebraicIndependent.ringHom_of_comp_eq (H : AlgebraicIndependent R x) AlgebraicIndependent S (g ∘ x) := by rw [algebraicIndependent_iff] at H ⊢ intro p hp - obtain ⟨q, rfl⟩ := map_surjective f hf p + obtain ⟨q, rfl⟩ := map_surjective (f : R →+* S) hf p rw [H q (hg (by rwa [map_zero, ← RingHom.coe_coe g, map_aeval, ← h, ← eval₂Hom_map_hom, ← aeval_eq_eval₂Hom])), map_zero] diff --git a/Mathlib/RingTheory/Ideal/Operations.lean b/Mathlib/RingTheory/Ideal/Operations.lean index 7218ec0a9688e..af540403ae860 100644 --- a/Mathlib/RingTheory/Ideal/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Operations.lean @@ -422,11 +422,23 @@ theorem prod_mem_prod {ι : Type*} {s : Finset ι} {I : ι → Ideal R} {x : ι theorem mul_le_right : I * J ≤ I := Ideal.mul_le.2 fun _ hr _ _ => I.mul_mem_right _ hr -@[simp] +#adaptation_note +/-- +On nightly-2024-11-12, we had to add `nolint simpNF` to the following lemma, +as otherwise we get a deterministic timeout in typeclass inference. +This should be investigated. +-/ +@[simp, nolint simpNF] theorem sup_mul_right_self : I ⊔ I * J = I := sup_eq_left.2 Ideal.mul_le_right -@[simp] +#adaptation_note +/-- +On nightly-2024-11-12, we had to add `nolint simpNF` to the following lemma, +as otherwise we get a deterministic timeout in typeclass inference. +This should be investigated. +-/ +@[simp, nolint simpNF] theorem mul_right_self_sup : I * J ⊔ I = I := sup_eq_right.2 Ideal.mul_le_right diff --git a/Mathlib/RingTheory/Jacobson.lean b/Mathlib/RingTheory/Jacobson.lean index 56997e889c66f..5fa8e8a0b7c09 100644 --- a/Mathlib/RingTheory/Jacobson.lean +++ b/Mathlib/RingTheory/Jacobson.lean @@ -702,7 +702,12 @@ lemma finite_of_finite_type_of_isJacobsonRing (R S : Type*) [CommRing R] [Field obtain ⟨ι, hι, f, hf⟩ := Algebra.FiniteType.iff_quotient_mvPolynomial'.mp ‹_› have : (algebraMap R S).IsIntegral := by rw [← f.comp_algebraMap] - exact MvPolynomial.comp_C_integral_of_surjective_of_isJacobsonRing f hf + #adaptation_note + /-- + After https://github.com/leanprover/lean4/pull/6024 + we needed to write `f.toRingHom` instead of just `f`, to avoid unification issues. + -/ + exact MvPolynomial.comp_C_integral_of_surjective_of_isJacobsonRing f.toRingHom hf have : Algebra.IsIntegral R S := Algebra.isIntegral_def.mpr this exact Algebra.IsIntegral.finite diff --git a/Mathlib/RingTheory/Multiplicity.lean b/Mathlib/RingTheory/Multiplicity.lean index c095227abd535..f6528327fbc78 100644 --- a/Mathlib/RingTheory/Multiplicity.lean +++ b/Mathlib/RingTheory/Multiplicity.lean @@ -414,7 +414,7 @@ theorem Nat.multiplicity_finite_iff {a b : ℕ} : Finite a b ↔ a ≠ 1 ∧ 0 < | 0 => ha rfl | 1 => ha1 rfl | b+2 => by omega - not_lt_of_ge (le_of_dvd (Nat.pos_of_ne_zero hb) (h b)) (lt_pow_self ha_gt_one b), + not_lt_of_ge (le_of_dvd (Nat.pos_of_ne_zero hb) (h b)) (b.lt_pow_self ha_gt_one), fun h => by cases h <;> simp [*]⟩ alias ⟨_, Dvd.multiplicity_pos⟩ := dvd_iff_multiplicity_pos diff --git a/Mathlib/RingTheory/Polynomial/UniqueFactorization.lean b/Mathlib/RingTheory/Polynomial/UniqueFactorization.lean index 6e24a23d71aee..7d5e251cc6781 100644 --- a/Mathlib/RingTheory/Polynomial/UniqueFactorization.lean +++ b/Mathlib/RingTheory/Polynomial/UniqueFactorization.lean @@ -128,7 +128,7 @@ instance (priority := 100) uniqueFactorizationMonoid : exact ⟨w.map (rename (↑)), fun b hb => let ⟨b', hb', he⟩ := Multiset.mem_map.1 hb - he ▸ (prime_rename_iff ↑s).2 (h b' hb'), + he ▸ (prime_rename_iff (σ := σ) ↑s).2 (h b' hb'), Units.map (@rename s σ D _ (↑)).toRingHom.toMonoidHom u, by erw [Multiset.prod_hom, ← map_mul, hw]⟩ diff --git a/Mathlib/RingTheory/WittVector/Frobenius.lean b/Mathlib/RingTheory/WittVector/Frobenius.lean index 79cc7194f5851..8ae9671d90379 100644 --- a/Mathlib/RingTheory/WittVector/Frobenius.lean +++ b/Mathlib/RingTheory/WittVector/Frobenius.lean @@ -124,7 +124,7 @@ theorem map_frobeniusPoly.key₂ {n i j : ℕ} (hi : i ≤ n) (hj : j < p ^ (n - tsub_add_cancel_of_le (le_tsub_of_add_le_right ((le_tsub_iff_left hi).mp h₁))] have hle : p ^ m ≤ j + 1 := h ▸ Nat.le_of_dvd j.succ_pos (pow_multiplicity_dvd _ _) exact ⟨(Nat.pow_le_pow_iff_right hp.1.one_lt).1 (hle.trans hj), - Nat.le_of_lt_succ ((Nat.lt_pow_self hp.1.one_lt m).trans_le hle)⟩ + Nat.le_of_lt_succ ((m.lt_pow_self hp.1.one_lt).trans_le hle)⟩ theorem map_frobeniusPoly (n : ℕ) : MvPolynomial.map (Int.castRingHom ℚ) (frobeniusPoly p n) = frobeniusPolyRat p n := by diff --git a/Mathlib/RingTheory/WittVector/InitTail.lean b/Mathlib/RingTheory/WittVector/InitTail.lean index a1c9c6d8601ac..0568831e3df4f 100644 --- a/Mathlib/RingTheory/WittVector/InitTail.lean +++ b/Mathlib/RingTheory/WittVector/InitTail.lean @@ -101,7 +101,7 @@ theorem select_add_select_not : ∀ x : 𝕎 R, select P x + select (fun i => ¬ refine fun m _ => mul_eq_mul_left_iff.mpr (Or.inl ?_) rw [ite_pow, zero_pow (pow_ne_zero _ hp.out.ne_zero)] by_cases Pm : P m - · rw [if_pos Pm, if_neg <| not_not_intro Pm, zero_pow Fin.size_pos'.ne', add_zero] + · rw [if_pos Pm, if_neg <| not_not_intro Pm, zero_pow Fin.pos'.ne', add_zero] · rwa [if_neg Pm, if_pos, zero_add] theorem coeff_add_of_disjoint (x y : 𝕎 R) (h : ∀ n, x.coeff n = 0 ∨ y.coeff n = 0) : diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index e0105b3587892..d4660d1f437be 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -3,7 +3,6 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Floris van Doorn -/ -import Mathlib.Algebra.Order.Ring.Nat import Mathlib.Data.Fintype.BigOperators import Mathlib.Data.Fintype.Powerset import Mathlib.Data.Nat.Cast.Order.Basic @@ -15,6 +14,8 @@ import Mathlib.Order.ConditionallyCompleteLattice.Indexed import Mathlib.Order.InitialSeg import Mathlib.Order.SuccPred.CompleteLinearOrder import Mathlib.SetTheory.Cardinal.SchroederBernstein +import Mathlib.Algebra.Order.GroupWithZero.Canonical +import Mathlib.Algebra.Order.Ring.Canonical /-! # Cardinal Numbers @@ -1778,12 +1779,12 @@ theorem mk_plift_false : #(PLift False) = 0 := mk_eq_zero _ @[simp] -theorem mk_vector (α : Type u) (n : ℕ) : #(Vector α n) = #α ^ n := +theorem mk_vector (α : Type u) (n : ℕ) : #(Mathlib.Vector α n) = #α ^ n := (mk_congr (Equiv.vectorEquivFin α n)).trans <| by simp theorem mk_list_eq_sum_pow (α : Type u) : #(List α) = sum fun n : ℕ => #α ^ n := calc - #(List α) = #(Σn, Vector α n) := mk_congr (Equiv.sigmaFiberEquiv List.length).symm + #(List α) = #(Σn, Mathlib.Vector α n) := mk_congr (Equiv.sigmaFiberEquiv List.length).symm _ = sum fun n : ℕ => #α ^ n := by simp theorem mk_quot_le {α : Type u} {r : α → α → Prop} : #(Quot r) ≤ #α := diff --git a/Mathlib/SetTheory/Game/PGame.lean b/Mathlib/SetTheory/Game/PGame.lean index 63640f6a65f95..4b18cfd466466 100644 --- a/Mathlib/SetTheory/Game/PGame.lean +++ b/Mathlib/SetTheory/Game/PGame.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Reid Barton, Mario Carneiro, Isabel Longbottom, Kim Morrison, Yuyang Zhao -/ import Mathlib.Algebra.Order.ZeroLEOne -import Mathlib.Data.List.InsertIdx import Mathlib.Logic.Relation import Mathlib.Logic.Small.Defs import Mathlib.Order.GameAdd diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index e7920e2f034c9..edf1824a8eb39 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -87,7 +87,7 @@ namespace WellOrder instance inhabited : Inhabited WellOrder := ⟨⟨PEmpty, _, inferInstanceAs (IsWellOrder PEmpty EmptyRelation)⟩⟩ -@[deprecated (since := "2024-10-24")] +@[deprecated "No deprecation message was provided." (since := "2024-10-24")] theorem eta (o : WellOrder) : mk o.α o.r o.wo = o := rfl end WellOrder @@ -137,12 +137,13 @@ instance inhabited : Inhabited Ordinal := instance one : One Ordinal := ⟨type <| @EmptyRelation PUnit⟩ -/-- Avoid using `Quotient.mk` to construct an `Ordinal` directly -/ -@[deprecated (since := "2024-10-24")] +@[deprecated "Avoid using `Quotient.mk` to construct an `Ordinal` directly." + (since := "2024-10-24")] theorem type_def' (w : WellOrder) : ⟦w⟧ = type w.r := rfl -/-- Avoid using `Quotient.mk` to construct an `Ordinal` directly -/ -@[deprecated (since := "2024-10-24")] + +@[deprecated "Avoid using `Quotient.mk` to construct an `Ordinal` directly." + (since := "2024-10-24")] theorem type_def (r) [wo : IsWellOrder α r] : (⟦⟨α, r, wo⟩⟧ : Ordinal) = type r := rfl @[simp] diff --git a/Mathlib/Tactic/Abel.lean b/Mathlib/Tactic/Abel.lean index 4002862603dbc..0326607cfdb06 100644 --- a/Mathlib/Tactic/Abel.lean +++ b/Mathlib/Tactic/Abel.lean @@ -436,14 +436,14 @@ The core of `abel_nf`, which rewrites the expression `e` into `abel` normal form -/ partial def abelNFCore (s : IO.Ref AtomM.State) (cfg : AbelNF.Config) (e : Expr) : MetaM Simp.Result := do - let ctx := { - simpTheorems := #[← Elab.Tactic.simpOnlyBuiltins.foldlM (·.addConst ·) {}] - congrTheorems := ← getSimpCongrTheorems } + let ctx ← Simp.mkContext + (simpTheorems := #[← Elab.Tactic.simpOnlyBuiltins.foldlM (·.addConst ·) {}]) + (congrTheorems := ← getSimpCongrTheorems) let simp ← match cfg.mode with | .raw => pure pure | .term => let thms := [``term_eq, ``termg_eq, ``add_zero, ``one_nsmul, ``one_zsmul, ``zsmul_zero] - let ctx' := { ctx with simpTheorems := #[← thms.foldlM (·.addConst ·) {:_}] } + let ctx' := ctx.setSimpTheorems #[← thms.foldlM (·.addConst ·) {:_}] pure fun r' : Simp.Result ↦ do r'.mkEqTrans (← Simp.main r'.expr ctx' (methods := ← Lean.Meta.Simp.mkDefaultMethods)).1 let rec diff --git a/Mathlib/Tactic/CC/Addition.lean b/Mathlib/Tactic/CC/Addition.lean index 402aa77e105b8..8548c1ad7b0c7 100644 --- a/Mathlib/Tactic/CC/Addition.lean +++ b/Mathlib/Tactic/CC/Addition.lean @@ -390,7 +390,7 @@ partial def mkCongrProofCore (lhs rhs : Expr) (heqProofs : Bool) : CCM Expr := d guard (kindsIt[0]! matches .eq) let some p ← getEqProof (lhsArgs[i]'hi.2) (rhsArgs[i]'(ha.symm ▸ hi.2)) | failure lemmaArgs := lemmaArgs.push p - kindsIt := kindsIt.eraseIdx 0 + kindsIt := kindsIt.eraseIdx! 0 let mut r := mkAppN specLemma.proof lemmaArgs if specLemma.heqResult && !heqProofs then r ← mkAppM ``eq_of_heq #[r] @@ -1748,7 +1748,7 @@ def propagateProjectionConstructor (p c : Expr) : CCM Unit := do unless ← pureIsDefEq (← inferType (pArgs[mkidx]'h)) (← inferType c) do return /- Create new projection application using c (e.g., `(x, y).fst`), and internalize it. The internalizer will add the new equality. -/ - let pArgs := pArgs.set ⟨mkidx, h⟩ c + let pArgs := pArgs.set mkidx c let newP := mkAppN pFn pArgs internalizeCore newP none else diff --git a/Mathlib/Tactic/CancelDenoms/Core.lean b/Mathlib/Tactic/CancelDenoms/Core.lean index b8ada93499ee9..db9ba05cb8083 100644 --- a/Mathlib/Tactic/CancelDenoms/Core.lean +++ b/Mathlib/Tactic/CancelDenoms/Core.lean @@ -231,7 +231,7 @@ def derive (e : Expr) : MetaM (ℕ × Expr) := do trace[CancelDenoms] "e = {e}" let eSimp ← simpOnlyNames (config := Simp.neutralConfig) deriveThms e trace[CancelDenoms] "e simplified = {eSimp.expr}" - let eSimpNormNum ← Mathlib.Meta.NormNum.deriveSimp {} false eSimp.expr + let eSimpNormNum ← Mathlib.Meta.NormNum.deriveSimp (← Simp.mkContext) false eSimp.expr trace[CancelDenoms] "e norm_num'd = {eSimpNormNum.expr}" let (n, t) := findCancelFactor eSimpNormNum.expr let ⟨u, tp, e⟩ ← inferTypeQ' eSimpNormNum.expr diff --git a/Mathlib/Tactic/CategoryTheory/BicategoryCoherence.lean b/Mathlib/Tactic/CategoryTheory/BicategoryCoherence.lean index 887a1d5d3c240..8aeea85868649 100644 --- a/Mathlib/Tactic/CategoryTheory/BicategoryCoherence.lean +++ b/Mathlib/Tactic/CategoryTheory/BicategoryCoherence.lean @@ -115,7 +115,7 @@ def bicategory_coherence (g : MVarId) : TermElabM Unit := g.withContext do let thms := [``BicategoricalCoherence.iso, ``Iso.trans, ``Iso.symm, ``Iso.refl, ``Bicategory.whiskerRightIso, ``Bicategory.whiskerLeftIso].foldl (·.addDeclToUnfoldCore ·) {} - let (ty, _) ← dsimp (← g.getType) { simpTheorems := #[thms] } + let (ty, _) ← dsimp (← g.getType) (← Simp.mkContext (simpTheorems := #[thms])) let some (_, lhs, rhs) := (← whnfR ty).eq? | exception g "Not an equation of morphisms." let lift_lhs ← mkLiftMap₂LiftExpr lhs let lift_rhs ← mkLiftMap₂LiftExpr rhs diff --git a/Mathlib/Tactic/CategoryTheory/Coherence.lean b/Mathlib/Tactic/CategoryTheory/Coherence.lean index e7e501d21c186..d136aaa42960b 100644 --- a/Mathlib/Tactic/CategoryTheory/Coherence.lean +++ b/Mathlib/Tactic/CategoryTheory/Coherence.lean @@ -124,7 +124,7 @@ def monoidal_coherence (g : MVarId) : TermElabM Unit := g.withContext do let thms := [``MonoidalCoherence.iso, ``Iso.trans, ``Iso.symm, ``Iso.refl, ``MonoidalCategory.whiskerRightIso, ``MonoidalCategory.whiskerLeftIso].foldl (·.addDeclToUnfoldCore ·) {} - let (ty, _) ← dsimp (← g.getType) { simpTheorems := #[thms] } + let (ty, _) ← dsimp (← g.getType) (← Simp.mkContext (simpTheorems := #[thms])) let some (_, lhs, rhs) := (← whnfR ty).eq? | exception g "Not an equation of morphisms." let projectMap_lhs ← mkProjectMapExpr lhs let projectMap_rhs ← mkProjectMapExpr rhs diff --git a/Mathlib/Tactic/CategoryTheory/Elementwise.lean b/Mathlib/Tactic/CategoryTheory/Elementwise.lean index 4a108c00b9927..656aab46bfe30 100644 --- a/Mathlib/Tactic/CategoryTheory/Elementwise.lean +++ b/Mathlib/Tactic/CategoryTheory/Elementwise.lean @@ -88,7 +88,7 @@ def elementwiseExpr (src : Name) (type pf : Expr) (simpSides := true) : lemmas, which can be caused by how applications are unfolded. \ Using elementwise is unnecessary." if simpSides then - let ctx := { ← Simp.Context.mkDefault with config.decide := false } + let ctx ← Simp.Context.mkDefault let (ty', eqPf'') ← simpEq (fun e => return (← simp e ctx).1) (← inferType eqPf') eqPf' -- check that it's not a simp-trivial equality: forallTelescope ty' fun _ ty' => do diff --git a/Mathlib/Tactic/DeprecateTo.lean b/Mathlib/Tactic/DeprecateTo.lean index 4f8d8522506e0..a374440b2c34e 100644 --- a/Mathlib/Tactic/DeprecateTo.lean +++ b/Mathlib/Tactic/DeprecateTo.lean @@ -107,7 +107,7 @@ Technically, the command also take an optional `String` argument to fill in the However, its use is mostly intended for debugging purposes, where having a variable date would make tests time-dependent. -/ -elab tk:"deprecate to " id:ident* dat:(ppSpace str ppSpace)? ppLine cmd:command : command => do +elab tk:"deprecate" "to" id:ident* dat:(ppSpace str ppSpace)? ppLine cmd:command : command => do let oldEnv ← getEnv try elabCommand cmd diff --git a/Mathlib/Tactic/DeriveTraversable.lean b/Mathlib/Tactic/DeriveTraversable.lean index 2ad019823e4bd..54a52f0b432bf 100644 --- a/Mathlib/Tactic/DeriveTraversable.lean +++ b/Mathlib/Tactic/DeriveTraversable.lean @@ -274,7 +274,7 @@ def deriveLawfulFunctor (m : MVarId) : TermElabM Unit := do if b then let hs ← getPropHyps s ← hs.foldlM (fun s f => f.getDecl >>= fun d => s.add (.fvar f) #[] d.toExpr) s - return { simpTheorems := #[s] } + Simp.mkContext (simpTheorems := #[s]) let .app (.app (.const ``LawfulFunctor _) F) _ ← m.getType >>= instantiateMVars | failure let some n := F.getAppFn.constName? | failure let [mcn, mim, mcm] ← m.applyConst ``LawfulFunctor.mk | failure @@ -435,7 +435,7 @@ def simpFunctorGoal (m : MVarId) (s : Simp.Context) (simprocs : Simp.SimprocsArr MetaM (Option (Array FVarId × MVarId) × Simp.Stats) := do let some e ← getSimpExtension? `functor_norm | failure let s' ← e.getTheorems - simpGoal m { s with simpTheorems := s.simpTheorems.push s' } simprocs discharge? simplifyTarget + simpGoal m (s.setSimpTheorems (s.simpTheorems.push s')) simprocs discharge? simplifyTarget fvarIdsToSimp stats /-- Run the following tactic: @@ -451,7 +451,7 @@ def traversableLawStarter (m : MVarId) (n : Name) (s : MetaM Simp.Context) (fun s n => s.addDeclToUnfold n) ({} : SimpTheorems) let (fi, m) ← m.intros m.withContext do - if let (some m, _) ← dsimpGoal m { simpTheorems := #[s'] } then + if let (some m, _) ← dsimpGoal m (← Simp.mkContext (simpTheorems := #[s'])) then let ma ← m.induction fi.back! (mkRecName n) ma.forM fun is => is.mvarId.withContext do @@ -467,9 +467,7 @@ def deriveLawfulTraversable (m : MVarId) : TermElabM Unit := do if b then let hs ← getPropHyps s ← hs.foldlM (fun s f => f.getDecl >>= fun d => s.add (.fvar f) #[] d.toExpr) s - pure <| - { config := { failIfUnchanged := false, unfoldPartialApp := true }, - simpTheorems := #[s] } + Simp.mkContext { failIfUnchanged := false, unfoldPartialApp := true } (simpTheorems := #[s]) let .app (.app (.const ``LawfulTraversable _) F) _ ← m.getType >>= instantiateMVars | failure let some n := F.getAppFn.constName? | failure let [mit, mct, mtmi, mn] ← m.applyConst ``LawfulTraversable.mk | failure diff --git a/Mathlib/Tactic/FBinop.lean b/Mathlib/Tactic/FBinop.lean index 3eb83fec8d8d6..8e58cb0849763 100644 --- a/Mathlib/Tactic/FBinop.lean +++ b/Mathlib/Tactic/FBinop.lean @@ -185,12 +185,12 @@ private def toExprCore (t : Tree) : TermElabM Expr := do | .term _ trees e => modifyInfoState (fun s => { s with trees := s.trees ++ trees }); return e | .binop ref f lhs rhs => - withRef ref <| withInfoContext' ref (mkInfo := mkTermInfo .anonymous ref) do + withRef ref <| withTermInfoContext' .anonymous ref do let lhs ← toExprCore lhs let mut rhs ← toExprCore rhs mkBinOp f lhs rhs | .macroExpansion macroName stx stx' nested => - withRef stx <| withInfoContext' stx (mkInfo := mkTermInfo macroName stx) do + withRef stx <| withTermInfoContext' macroName stx do withMacroExpansion stx stx' do toExprCore nested diff --git a/Mathlib/Tactic/FieldSimp.lean b/Mathlib/Tactic/FieldSimp.lean index b401457965534..a3e7f13ab83a1 100644 --- a/Mathlib/Tactic/FieldSimp.lean +++ b/Mathlib/Tactic/FieldSimp.lean @@ -61,26 +61,27 @@ partial def discharge (prop : Expr) : SimpM (Option Expr) := if let some pf := pf? then return some pf -- Discharge strategy 4: Use the simplifier - let ctx ← readThe Simp.Context - let stats : Simp.Stats := { (← get) with } - - -- Porting note: mathlib3's analogous field_simp discharger `field_simp.ne_zero` - -- does not explicitly call `simp` recursively like this. It's unclear to me - -- whether this is because - -- 1) Lean 3 simp dischargers automatically call `simp` recursively. (Do they?), - -- 2) mathlib3 norm_num1 is able to handle any needed discharging, or - -- 3) some other reason? - let ⟨simpResult, stats'⟩ ← - simp prop { ctx with dischargeDepth := ctx.dischargeDepth + 1 } #[(← Simp.getSimprocs)] - discharge stats - set { (← get) with usedTheorems := stats'.usedTheorems, diag := stats'.diag } - if simpResult.expr.isConstOf ``True then - try - return some (← mkOfEqTrue (← simpResult.getProof)) - catch _ => + Simp.withIncDischargeDepth do + let ctx ← readThe Simp.Context + let stats : Simp.Stats := { (← get) with } + + -- Porting note: mathlib3's analogous field_simp discharger `field_simp.ne_zero` + -- does not explicitly call `simp` recursively like this. It's unclear to me + -- whether this is because + -- 1) Lean 3 simp dischargers automatically call `simp` recursively. (Do they?), + -- 2) mathlib3 norm_num1 is able to handle any needed discharging, or + -- 3) some other reason? + let ⟨simpResult, stats'⟩ ← + simp prop ctx #[(← Simp.getSimprocs)] + discharge stats + set { (← get) with usedTheorems := stats'.usedTheorems, diag := stats'.diag } + if simpResult.expr.isConstOf ``True then + try + return some (← mkOfEqTrue (← simpResult.getProof)) + catch _ => + return none + else return none - else - return none @[inherit_doc discharge] elab "field_simp_discharge" : tactic => wrapSimpDischarger Mathlib.Tactic.FieldSimp.discharge @@ -175,11 +176,10 @@ elab_rules : tactic let some ext ← getSimpExtension? `field_simps | throwError "field_simps not found" let thms ← ext.getTheorems - let ctx : Simp.Context := { - simpTheorems := #[thms, thms0] - congrTheorems := ← getSimpCongrTheorems - config := cfg - } + let ctx ← Simp.mkContext cfg + (simpTheorems := #[thms, thms0]) + (congrTheorems := ← getSimpCongrTheorems) + let mut r ← elabSimpArgs (sa.getD ⟨.missing⟩) ctx (simprocs := {}) (eraseLocal := false) .simp if r.starArg then r ← do @@ -189,7 +189,7 @@ elab_rules : tactic for h in hs do unless simpTheorems.isErased (.fvar h) do simpTheorems ← simpTheorems.addTheorem (.fvar h) (← h.getDecl).toExpr - let ctx := { ctx with simpTheorems } + let ctx := ctx.setSimpTheorems simpTheorems pure { ctx, simprocs := {} } _ ← simpLocation r.ctx {} dis loc diff --git a/Mathlib/Tactic/FunProp/Attr.lean b/Mathlib/Tactic/FunProp/Attr.lean index 55a876dcd0943..5566c0f969179 100644 --- a/Mathlib/Tactic/FunProp/Attr.lean +++ b/Mathlib/Tactic/FunProp/Attr.lean @@ -27,7 +27,6 @@ initialize funPropAttr : Unit ← add := fun declName _stx attrKind => discard <| MetaM.run do let info ← getConstInfo declName - forallTelescope info.type fun _ b => do if b.isProp then addFunPropDecl declName diff --git a/Mathlib/Tactic/FunProp/Core.lean b/Mathlib/Tactic/FunProp/Core.lean index 9f84b5b2c1371..105bd3beea9c0 100644 --- a/Mathlib/Tactic/FunProp/Core.lean +++ b/Mathlib/Tactic/FunProp/Core.lean @@ -324,9 +324,7 @@ def applyMorRules (funPropDecl : FunPropDecl) (e : Expr) (fData : FunctionData) applyCompRule funPropDecl e f g funProp | .exact => - let ext := morTheoremsExt.getState (← getEnv) - let candidates ← ext.theorems.getMatchWithScore e false { iota := false, zeta := false } - let candidates := candidates.map (·.1) |>.flatten + let candidates ← getMorphismTheorems e trace[Meta.Tactic.fun_prop] "candidate morphism theorems: {← candidates.mapM fun c => ppOrigin (.decl c.thmName)}" @@ -343,9 +341,7 @@ def applyTransitionRules (e : Expr) (funProp : Expr → FunPropM (Option Result) FunPropM (Option Result) := do withIncreasedTransitionDepth do - let ext := transitionTheoremsExt.getState (← getEnv) - let candidates ← ext.theorems.getMatchWithScore e false { iota := false, zeta := false } - let candidates := candidates.map (·.1) |>.flatten + let candidates ← getTransitionTheorems e trace[Meta.Tactic.fun_prop] "candidate transition theorems: {← candidates.mapM fun c => ppOrigin (.decl c.thmName)}" @@ -433,7 +429,7 @@ def getLocalTheorems (funPropDecl : FunPropDecl) (funOrigin : Origin) let .some (decl,f) ← getFunProp? b | return none unless decl.funPropName = funPropDecl.funPropName do return none - let .data fData ← getFunctionData? f (← unfoldNamePred) {zeta := false, zetaDelta := false} + let .data fData ← getFunctionData? f (← unfoldNamePred) | return none unless (fData.getFnOrigin == funOrigin) do return none @@ -672,7 +668,7 @@ mutual let e' := e.setArg funPropDecl.funArgId b funProp (← mkLambdaFVars xs e') - match ← getFunctionData? f (← unfoldNamePred) {zeta := false, zetaDelta := false} with + match ← getFunctionData? f (← unfoldNamePred) with | .letE f => trace[Debug.Meta.Tactic.fun_prop] "let case on {← ppExpr f}" let e := e.setArg funPropDecl.funArgId f -- update e with reduced f diff --git a/Mathlib/Tactic/FunProp/Decl.lean b/Mathlib/Tactic/FunProp/Decl.lean index d91b7a71fec56..eecfb35cb2fdd 100644 --- a/Mathlib/Tactic/FunProp/Decl.lean +++ b/Mathlib/Tactic/FunProp/Decl.lean @@ -61,7 +61,7 @@ def addFunPropDecl (declName : Name) : MetaM Unit := do let lvls := info.levelParams.map (fun l => Level.param l) let e := mkAppN (.const declName lvls) xs - let path ← DiscrTree.mkPath e {} + let path ← DiscrTree.mkPath e -- find the argument position of the function `f` in `P f` let mut .some funArgId ← (xs.zip bi).findIdxM? fun (x,bi) => do @@ -88,7 +88,7 @@ the function it talks about. -/ def getFunProp? (e : Expr) : MetaM (Option (FunPropDecl × Expr)) := do let ext := funPropDeclsExt.getState (← getEnv) - let decls ← ext.decls.getMatch e {} + let decls ← ext.decls.getMatch e (← read) if decls.size = 0 then return none diff --git a/Mathlib/Tactic/FunProp/FunctionData.lean b/Mathlib/Tactic/FunProp/FunctionData.lean index c5f3356767d0f..8c36005b51303 100644 --- a/Mathlib/Tactic/FunProp/FunctionData.lean +++ b/Mathlib/Tactic/FunProp/FunctionData.lean @@ -117,8 +117,9 @@ def MaybeFunctionData.get (fData : MaybeFunctionData) : MetaM Expr := /-- Get `FunctionData` for `f`. -/ def getFunctionData? (f : Expr) - (unfoldPred : Name → Bool := fun _ => false) (cfg : WhnfCoreConfig := {}) : + (unfoldPred : Name → Bool := fun _ => false) : MetaM MaybeFunctionData := do + withConfig (fun cfg => { cfg with zeta := false, zetaDelta := false }) do let unfold := fun e : Expr => do if let .some n := e.getAppFn'.constName? then @@ -130,7 +131,7 @@ def getFunctionData? (f : Expr) | throwError m!"fun_prop bug: function expected, got `{f} : {← inferType f}, \ type ctor {(← inferType f).ctorName}" withLocalDeclD xName xType fun x => do - let fx' := (← Mor.whnfPred (f.beta #[x]).eta unfold cfg) |> headBetaThroughLet + let fx' := (← Mor.whnfPred (f.beta #[x]).eta unfold) |> headBetaThroughLet let f' ← mkLambdaFVars #[x] fx' match fx' with | .letE .. => return .letE f' diff --git a/Mathlib/Tactic/FunProp/Mor.lean b/Mathlib/Tactic/FunProp/Mor.lean index 1b27bf7755713..b7cf3ec7c15e2 100644 --- a/Mathlib/Tactic/FunProp/Mor.lean +++ b/Mathlib/Tactic/FunProp/Mor.lean @@ -63,14 +63,14 @@ can specify which when to unfold definitions. For example calling this on `coe (f a) b` will put `f` in weak normal head form instead of `coe`. -/ -partial def whnfPred (e : Expr) (pred : Expr → MetaM Bool) (cfg : WhnfCoreConfig := {}) : +partial def whnfPred (e : Expr) (pred : Expr → MetaM Bool) : MetaM Expr := do whnfEasyCases e fun e => do - let e ← whnfCore e cfg + let e ← whnfCore e if let .some ⟨coe,f,x⟩ ← isMorApp? e then - let f ← whnfPred f pred cfg - if cfg.zeta then + let f ← whnfPred f pred + if (← getConfig).zeta then return (coe.app f).app x else return ← letTelescope f fun xs f' => @@ -78,7 +78,7 @@ partial def whnfPred (e : Expr) (pred : Expr → MetaM Bool) (cfg : WhnfCoreConf if (← pred e) then match (← unfoldDefinition? e) with - | some e => whnfPred e pred cfg + | some e => whnfPred e pred | none => return e else return e @@ -88,8 +88,8 @@ Weak normal head form of an expression involving morphism applications. For example calling this on `coe (f a) b` will put `f` in weak normal head form instead of `coe`. -/ -def whnf (e : Expr) (cfg : WhnfCoreConfig := {}) : MetaM Expr := - whnfPred e (fun _ => return false) cfg +def whnf (e : Expr) : MetaM Expr := + whnfPred e (fun _ => return false) /-- Argument of morphism application that stores corresponding coercion if necessary -/ diff --git a/Mathlib/Tactic/FunProp/RefinedDiscrTree.lean b/Mathlib/Tactic/FunProp/RefinedDiscrTree.lean index 46b53085e8011..439560f236c24 100644 --- a/Mathlib/Tactic/FunProp/RefinedDiscrTree.lean +++ b/Mathlib/Tactic/FunProp/RefinedDiscrTree.lean @@ -438,23 +438,23 @@ where | _ => failure /-- Reduction procedure for the `RefinedDiscrTree` indexing. -/ -partial def reduce (e : Expr) (config : WhnfCoreConfig) : MetaM Expr := do - let e ← whnfCore e config +partial def reduce (e : Expr) : MetaM Expr := do + let e ← whnfCore e match (← unfoldDefinition? e) with - | some e => reduce e config + | some e => reduce e | none => match e.etaExpandedStrict? with - | some e => reduce e config + | some e => reduce e | none => return e /-- Repeatedly apply reduce while stripping lambda binders and introducing their variables -/ @[specialize] partial def lambdaTelescopeReduce {m} [Monad m] [MonadLiftT MetaM m] [MonadControlT MetaM m] - [Nonempty α] (e : Expr) (fvars : List FVarId) (config : WhnfCoreConfig) + [Nonempty α] (e : Expr) (fvars : List FVarId) (k : Expr → List FVarId → m α) : m α := do - match ← reduce e config with + match ← reduce e with | .lam n d b bi => withLocalDecl n bi d fun fvar => - lambdaTelescopeReduce (b.instantiate1 fvar) (fvar.fvarId! :: fvars) config k + lambdaTelescopeReduce (b.instantiate1 fvar) (fvar.fvarId! :: fvars) k | e => k e fvars @@ -492,7 +492,6 @@ private structure Context where bvars : List FVarId := [] /-- Variables that come from a lambda that has been removed via η-reduction. -/ forbiddenVars : List FVarId := [] - config : WhnfCoreConfig fvarInContext : FVarId → Bool /-- Return for each argument whether it should be ignored. -/ @@ -633,7 +632,7 @@ private def withLams {m} [Monad m] [MonadWithReader Context m] /-- Return the encoding of `e` as a `DTExpr`. If `root = false`, then `e` is a strict sub expression of the original expression. -/ partial def mkDTExprAux (e : Expr) (root : Bool) : ReaderT Context MetaM DTExpr := do - lambdaTelescopeReduce e [] (← read).config fun e lambdas => + lambdaTelescopeReduce e [] fun e lambdas => e.withApp fun fn args => do let argDTExpr (arg : Expr) (ignore : Bool) : ReaderT Context MetaM DTExpr := @@ -755,7 +754,7 @@ def cacheEtaPossibilities (e original : Expr) (lambdas : List FVarId) /-- Return all encodings of `e` as a `DTExpr`, taking possible η-reductions into account. If `root = false`, then `e` is a strict sub expression of the original expression. -/ partial def mkDTExprsAux (original : Expr) (root : Bool) : M DTExpr := do - lambdaTelescopeReduce original [] (← read).config fun e lambdas => do + lambdaTelescopeReduce original [] fun e lambdas => do if !root then if let .const n _ := e.getAppFn then @@ -849,16 +848,16 @@ Warning: to account for potential η-reductions of `e`, use `mkDTExprs` instead. The argument `fvarInContext` allows you to specify which free variables in `e` will still be in the context when the `RefinedDiscrTree` is being used for lookup. It should return true only if the `RefinedDiscrTree` is built and used locally. -/ -def mkDTExpr (e : Expr) (config : WhnfCoreConfig) +def mkDTExpr (e : Expr) (fvarInContext : FVarId → Bool := fun _ => false) : MetaM DTExpr := - withReducible do (MkDTExpr.mkDTExprAux e true |>.run {config, fvarInContext}) + withReducible do (MkDTExpr.mkDTExprAux e true |>.run {fvarInContext}) /-- Similar to `mkDTExpr`. Return all encodings of `e` as a `DTExpr`, taking potential further η-reductions into account. -/ -def mkDTExprs (e : Expr) (config : WhnfCoreConfig) (onlySpecific : Bool) +def mkDTExprs (e : Expr) (onlySpecific : Bool) (fvarInContext : FVarId → Bool := fun _ => false) : MetaM (List DTExpr) := withReducible do - let es ← (MkDTExpr.mkDTExprsAux e true).run' {} |>.run {config, fvarInContext} + let es ← (MkDTExpr.mkDTExprsAux e true).run' {} |>.run {fvarInContext} return if onlySpecific then es.filter (·.isSpecific) else es @@ -875,7 +874,7 @@ where loop (i : Nat) : Array α := if h : i < vs.size then if v == vs[i] then - vs.set ⟨i,h⟩ v + vs.set i v else loop (i+1) else @@ -932,18 +931,18 @@ It should return true only if the `RefinedDiscrTree` is built and used locally. if `onlySpecific := true`, then we filter out the patterns `*` and `Eq * * *`. -/ def insert [BEq α] (d : RefinedDiscrTree α) (e : Expr) (v : α) - (onlySpecific : Bool := true) (config : WhnfCoreConfig := {}) - (fvarInContext : FVarId → Bool := fun _ => false) : MetaM (RefinedDiscrTree α) := do - let keys ← mkDTExprs e config onlySpecific fvarInContext + (onlySpecific : Bool := true) (fvarInContext : FVarId → Bool := fun _ => false) : + MetaM (RefinedDiscrTree α) := do + let keys ← mkDTExprs e onlySpecific fvarInContext return keys.foldl (insertDTExpr · · v) d /-- Insert the value `vLhs` at index `lhs`, and if `rhs` is indexed differently, then also insert the value `vRhs` at index `rhs`. -/ def insertEqn [BEq α] (d : RefinedDiscrTree α) (lhs rhs : Expr) (vLhs vRhs : α) - (onlySpecific : Bool := true) (config : WhnfCoreConfig := {}) - (fvarInContext : FVarId → Bool := fun _ => false) : MetaM (RefinedDiscrTree α) := do - let keysLhs ← mkDTExprs lhs config onlySpecific fvarInContext - let keysRhs ← mkDTExprs rhs config onlySpecific fvarInContext + (onlySpecific : Bool := true) (fvarInContext : FVarId → Bool := fun _ => false) : + MetaM (RefinedDiscrTree α) := do + let keysLhs ← mkDTExprs lhs onlySpecific fvarInContext + let keysRhs ← mkDTExprs rhs onlySpecific fvarInContext let d := keysLhs.foldl (insertDTExpr · · vLhs) d if @List.beq _ ⟨DTExpr.eqv⟩ keysLhs keysRhs then return d @@ -967,7 +966,6 @@ def findKey (children : Array (Key × Trie α)) (k : Key) : Option (Trie α) := private structure Context where unify : Bool - config : WhnfCoreConfig private structure State where /-- Score representing how good the match is. -/ @@ -981,9 +979,9 @@ private structure State where private abbrev M := ReaderT Context <| StateListM State /-- Return all values from `x` in an array, together with their scores. -/ -private def M.run (unify : Bool) (config : WhnfCoreConfig) (x : M (Trie α)) : +private def M.run (unify : Bool) (x : M (Trie α)) : Array (Array α × Nat) := - ((x.run { unify, config }).run {}).toArray.map (fun (t, s) => (t.values!, s.score)) + ((x.run { unify }).run {}).toArray.map (fun (t, s) => (t.values!, s.score)) /-- Increment the score by `n`. -/ private def incrementScore (n : Nat) : M Unit := @@ -1076,7 +1074,7 @@ mutual end private partial def getMatchWithScoreAux (d : RefinedDiscrTree α) (e : DTExpr) (unify : Bool) - (config : WhnfCoreConfig) (allowRootStar : Bool := false) : Array (Array α × Nat) := (do + (allowRootStar : Bool := false) : Array (Array α × Nat) := (do if e matches .star _ then guard allowRootStar d.root.foldl (init := failure) fun x k c => (do @@ -1090,7 +1088,7 @@ private partial def getMatchWithScoreAux (d : RefinedDiscrTree α) (e : DTExpr) guard allowRootStar let some c := d.root.find? (.star 0) | failure return c - ).run unify config + ).run unify end GetUnify @@ -1106,22 +1104,22 @@ This is for when you don't want to instantiate metavariables in `e`. If `allowRootStar := false`, then we don't allow `e` or the matched key in `d` to be a star pattern. -/ def getMatchWithScore (d : RefinedDiscrTree α) (e : Expr) (unify : Bool) - (config : WhnfCoreConfig) (allowRootStar : Bool := false) : MetaM (Array (Array α × Nat)) := do - let e ← mkDTExpr e config - let result := GetUnify.getMatchWithScoreAux d e unify config allowRootStar + (allowRootStar : Bool := false) : MetaM (Array (Array α × Nat)) := do + let e ← mkDTExpr e + let result := GetUnify.getMatchWithScoreAux d e unify allowRootStar return result.qsort (·.2 > ·.2) /-- Similar to `getMatchWithScore`, but also returns matches with prefixes of `e`. We store the score, followed by the number of ignored arguments. -/ partial def getMatchWithScoreWithExtra (d : RefinedDiscrTree α) (e : Expr) (unify : Bool) - (config : WhnfCoreConfig) (allowRootStar : Bool := false) : + (allowRootStar : Bool := false) : MetaM (Array (Array α × Nat × Nat)) := do let result ← go e 0 return result.qsort (·.2.1 > ·.2.1) where /-- go -/ go (e : Expr) (numIgnored : Nat) : MetaM (Array (Array α × Nat × Nat)) := do - let result ← getMatchWithScore d e unify config allowRootStar + let result ← getMatchWithScore d e unify allowRootStar let result := result.map fun (a, b) => (a, b, numIgnored) match e with | .app e _ => return (← go e (numIgnored + 1)) ++ result diff --git a/Mathlib/Tactic/FunProp/Theorems.lean b/Mathlib/Tactic/FunProp/Theorems.lean index b0c184aedc988..464ce021ba426 100644 --- a/Mathlib/Tactic/FunProp/Theorems.lean +++ b/Mathlib/Tactic/FunProp/Theorems.lean @@ -226,16 +226,16 @@ structure GeneralTheorem where def GeneralTheorem.getProof (thm : GeneralTheorem) : MetaM Expr := do mkConstWithFreshMVarLevels thm.thmName -/-- -/ +/-- Structure holding transition or morphism theorems for `fun_prop` tactic. -/ structure GeneralTheorems where - /-- -/ + /-- Discrimination tree indexing theorems. -/ theorems : RefinedDiscrTree GeneralTheorem := {} deriving Inhabited -/-- -/ +/-- Extendions for transition or morphism theorems -/ abbrev GeneralTheoremsExt := SimpleScopedEnvExtension GeneralTheorem GeneralTheorems -/-- -/ +/-- Environment extension for transition theorems. -/ initialize transitionTheoremsExt : GeneralTheoremsExt ← registerSimpleScopedEnvExtension { name := by exact decl_name% @@ -244,7 +244,18 @@ initialize transitionTheoremsExt : GeneralTheoremsExt ← {d with theorems := e.keys.foldl (RefinedDiscrTree.insertDTExpr · · e) d.theorems} } -/-- -/ +/-- Get transition theorems applicable to `e`. + +For example calling on `e` equal to `Continuous f` might return theorems implying continuity +from linearity over finite dimensional spaces or differentiability. -/ +def getTransitionTheorems (e : Expr) : FunPropM (Array GeneralTheorem) := do + let ext := transitionTheoremsExt.getState (← getEnv) + let candidates ← withConfig (fun cfg => { cfg with iota := false, zeta := false }) <| + ext.theorems.getMatchWithScore e false + let candidates := candidates.map (·.1) |>.flatten + return candidates + +/-- Environment extension for morphism theorems. -/ initialize morTheoremsExt : GeneralTheoremsExt ← registerSimpleScopedEnvExtension { name := by exact decl_name% @@ -254,6 +265,17 @@ initialize morTheoremsExt : GeneralTheoremsExt ← } +/-- Get morphism theorems applicable to `e`. + +For example calling on `e` equal to `Continuous f` for `f : X→L[ℝ] Y` would return theorem +infering continuity from the bundled morphism. -/ +def getMorphismTheorems (e : Expr) : FunPropM (Array GeneralTheorem) := do + let ext := morTheoremsExt.getState (← getEnv) + let candidates ← withConfig (fun cfg => { cfg with iota := false, zeta := false }) <| + ext.theorems.getMatchWithScore e false + let candidates := candidates.map (·.1) |>.flatten + return candidates + -------------------------------------------------------------------------------- @@ -301,13 +323,11 @@ type of theorem it is. -/ def getTheoremFromConst (declName : Name) (prio : Nat := eval_prio default) : MetaM Theorem := do let info ← getConstInfo declName forallTelescope info.type fun xs b => do - let .some (decl,f) ← getFunProp? b | throwError "unrecognized function property `{← ppExpr b}`" let funPropName := decl.funPropName - - let fData? ← getFunctionData? f defaultUnfoldPred {zeta := false} - + let fData? ← + withConfig (fun cfg => { cfg with zeta := false}) <| getFunctionData? f defaultUnfoldPred if let .some thmArgs ← detectLambdaTheoremArgs (← fData?.get) xs then return .lam { funPropName := funPropName @@ -338,7 +358,7 @@ def getTheoremFromConst (declName : Name) (prio : Nat := eval_prio default) : Me } | .fvar .. => let (_,_,b') ← forallMetaTelescope info.type - let keys := ← RefinedDiscrTree.mkDTExprs b' {} false + let keys := ← RefinedDiscrTree.mkDTExprs b' false let thm : GeneralTheorem := { funPropName := funPropName thmName := declName diff --git a/Mathlib/Tactic/FunProp/ToBatteries.lean b/Mathlib/Tactic/FunProp/ToBatteries.lean index f3affc6579c3d..f11a592a4c1ad 100644 --- a/Mathlib/Tactic/FunProp/ToBatteries.lean +++ b/Mathlib/Tactic/FunProp/ToBatteries.lean @@ -35,7 +35,7 @@ def isOrderedSubsetOf {α} [Inhabited α] [DecidableEq α] (a b : Array α) : Bo private def letTelescopeImpl {α} (e : Expr) (k : Array Expr → Expr → MetaM α) : MetaM α := lambdaLetTelescope e fun xs b ↦ do - if let .some i ← xs.findIdxM? (fun x ↦ do pure ¬(← x.fvarId!.isLetVar)) then + if let .some i ← xs.findIdxM? (fun x ↦ do pure !(← x.fvarId!.isLetVar)) then k xs[0:i] (← mkLambdaFVars xs[i:] b) else k xs b diff --git a/Mathlib/Tactic/Linarith/Frontend.lean b/Mathlib/Tactic/Linarith/Frontend.lean index 04b1e07d1ac97..a63112c459aff 100644 --- a/Mathlib/Tactic/Linarith/Frontend.lean +++ b/Mathlib/Tactic/Linarith/Frontend.lean @@ -277,7 +277,7 @@ def runLinarith (cfg : LinarithConfig) (prefType : Option Expr) (g : MVarId) if let some t := prefType then let (i, vs) ← hyp_set.find t proveFalseByLinarith cfg.transparency cfg.oracle cfg.discharger g vs <|> - findLinarithContradiction cfg g ((hyp_set.eraseIdx i).toList.map (·.2)) + findLinarithContradiction cfg g ((hyp_set.eraseIdxIfInBounds i).toList.map (·.2)) else findLinarithContradiction cfg g (hyp_set.toList.map (·.2)) let mut preprocessors := cfg.preprocessors if cfg.splitNe then diff --git a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/Datatypes.lean b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/Datatypes.lean index 8a5a9488ae1a9..2f6eb09ce7de1 100644 --- a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/Datatypes.lean +++ b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/Datatypes.lean @@ -69,7 +69,7 @@ instance : UsableInSimplexAlgorithm DenseMatrix where for ⟨i, j, v⟩ in vals do data := data.modify i fun row => row.set! j v return ⟨data⟩ - swapRows mat i j := ⟨mat.data.swap! i j⟩ + swapRows mat i j := ⟨mat.data.swapIfInBounds i j⟩ subtractRow mat i j coef := let newData : Array (Array Rat) := mat.data.modify j fun row => row.zipWith mat.data[i]! fun x y => x - coef * y @@ -101,7 +101,7 @@ instance : UsableInSimplexAlgorithm SparseMatrix where if v != 0 then data := data.modify i fun row => row.insert j v return ⟨data⟩ - swapRows mat i j := ⟨mat.data.swap! i j⟩ + swapRows mat i j := ⟨mat.data.swapIfInBounds i j⟩ subtractRow mat i j coef := let newData := mat.data.modify j fun row => mat.data[i]!.fold (fun cur k val => diff --git a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/SimplexAlgorithm.lean b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/SimplexAlgorithm.lean index 9858e94c23f14..4ecfb36c3cab1 100644 --- a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/SimplexAlgorithm.lean +++ b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/SimplexAlgorithm.lean @@ -48,8 +48,8 @@ def doPivotOperation (exitIdx enterIdx : Nat) : SimplexAlgorithmM matType Unit : let newBasic := s.basic.set! exitIdx s.free[enterIdx]! let newFree := s.free.set! enterIdx s.basic[exitIdx]! - have hb : newBasic.size = s.basic.size := by apply Array.size_setD - have hf : newFree.size = s.free.size := by apply Array.size_setD + have hb : newBasic.size = s.basic.size := by apply Array.size_setIfInBounds + have hf : newFree.size = s.free.size := by apply Array.size_setIfInBounds return (⟨newBasic, newFree, hb ▸ hf ▸ mat⟩ : Tableau matType) @@ -60,7 +60,7 @@ nonnegative. def checkSuccess : SimplexAlgorithmM matType Bool := do let lastIdx := (← get).free.size - 1 return (← get).mat[(0, lastIdx)]! > 0 && - (← Nat.allM (← get).basic.size (fun i => do return (← get).mat[(i, lastIdx)]! >= 0)) + (← (← get).basic.size.allM (fun i _ => do return (← get).mat[(i, lastIdx)]! ≥ 0)) /-- Chooses an entering variable: among the variables with a positive coefficient in the objective diff --git a/Mathlib/Tactic/Linarith/Preprocessing.lean b/Mathlib/Tactic/Linarith/Preprocessing.lean index eda221c8c33d1..348309a7774af 100644 --- a/Mathlib/Tactic/Linarith/Preprocessing.lean +++ b/Mathlib/Tactic/Linarith/Preprocessing.lean @@ -30,7 +30,7 @@ namespace Linarith /-! ### Preprocessing -/ -open Lean hiding Rat +open Lean open Elab Tactic Meta open Qq open Mathlib diff --git a/Mathlib/Tactic/LinearCombination'.lean b/Mathlib/Tactic/LinearCombination'.lean index c7965a7998ee9..d322231af38f8 100644 --- a/Mathlib/Tactic/LinearCombination'.lean +++ b/Mathlib/Tactic/LinearCombination'.lean @@ -36,7 +36,7 @@ implementation, but this version is provided for backward-compatibility. -/ namespace Mathlib.Tactic.LinearCombination' -open Lean hiding Rat +open Lean open Elab Meta Term variable {α : Type*} {a a' a₁ a₂ b b' b₁ b₂ c : α} diff --git a/Mathlib/Tactic/LinearCombination.lean b/Mathlib/Tactic/LinearCombination.lean index d5f5a2b863b2e..88ba7708252df 100644 --- a/Mathlib/Tactic/LinearCombination.lean +++ b/Mathlib/Tactic/LinearCombination.lean @@ -36,7 +36,7 @@ Lastly, calls a normalization tactic on this target. -/ namespace Mathlib.Tactic.LinearCombination -open Lean hiding Rat +open Lean open Elab Meta Term Mathlib Ineq /-- Result of `expandLinearCombo`, either an equality/inequality proof or a value. -/ diff --git a/Mathlib/Tactic/Linter/PPRoundtrip.lean b/Mathlib/Tactic/Linter/PPRoundtrip.lean index 9f8e258ecd46f..df238e0927ee3 100644 --- a/Mathlib/Tactic/Linter/PPRoundtrip.lean +++ b/Mathlib/Tactic/Linter/PPRoundtrip.lean @@ -60,7 +60,7 @@ def polishSource (s : String) : String × Array Nat := let preWS := split.foldl (init := #[]) fun p q => let txt := q.trimLeft.length (p.push (q.length - txt)).push txt - let preWS := preWS.eraseIdx 0 + let preWS := preWS.eraseIdxIfInBounds 0 let s := (split.map .trimLeft).filter (· != "") (" ".intercalate (s.filter (!·.isEmpty)), preWS) diff --git a/Mathlib/Tactic/MinImports.lean b/Mathlib/Tactic/MinImports.lean index 05f758440185c..c9544b2b3bae9 100644 --- a/Mathlib/Tactic/MinImports.lean +++ b/Mathlib/Tactic/MinImports.lean @@ -215,10 +215,10 @@ def minImpsCore (stx id : Syntax) : CommandElabM Unit := do /-- `#min_imports in cmd` scans the syntax `cmd` and the declaration obtained by elaborating `cmd` to find a collection of minimal imports that should be sufficient for `cmd` to work. -/ -syntax (name := minImpsStx) "#min_imports in" command : command +syntax (name := minImpsStx) "#min_imports" "in" command : command @[inherit_doc minImpsStx] -syntax "#min_imports in" term : command +syntax "#min_imports" "in" term : command elab_rules : command | `(#min_imports in $cmd:command) => do diff --git a/Mathlib/Tactic/Module.lean b/Mathlib/Tactic/Module.lean index e06d5ce2088c2..14274913e2c77 100644 --- a/Mathlib/Tactic/Module.lean +++ b/Mathlib/Tactic/Module.lean @@ -573,10 +573,7 @@ def postprocess (mvarId : MVarId) : MetaM MVarId := do let ⟨levelParams, _, proof⟩ ← abstractMVars (mkConst thm) thms ← thms.add (.stx (← mkFreshId) Syntax.missing) levelParams proof -- now run `simp` with these lemmas, and (importantly) *no* simprocs - let ctx : Simp.Context := { - config := { failIfUnchanged := false } - simpTheorems := #[thms] - } + let ctx ← Simp.mkContext { failIfUnchanged := false } (simpTheorems := #[thms]) let (some r, _) ← simpTarget mvarId ctx (simprocs := #[]) | throwError "internal error in match_scalars tactic: postprocessing should not close goals" return r diff --git a/Mathlib/Tactic/MoveAdd.lean b/Mathlib/Tactic/MoveAdd.lean index 4cee13436d339..12f7051a82e2c 100644 --- a/Mathlib/Tactic/MoveAdd.lean +++ b/Mathlib/Tactic/MoveAdd.lean @@ -195,7 +195,8 @@ def reorderUsing (toReorder : List α) (instructions : List (α × Bool)) : List let uToReorder := (uniquify toReorder).toArray let reorder := uToReorder.qsort fun x y => match uInstructions.find? (Prod.fst · == x), uInstructions.find? (Prod.fst · == y) with - | none, none => (uToReorder.getIdx? x).get! ≤ (uToReorder.getIdx? y).get! + | none, none => + ((uToReorder.indexOf? x).map Fin.val).get! ≤ ((uToReorder.indexOf? y).map Fin.val).get! | _, _ => weight uInstructions x ≤ weight uInstructions y (reorder.map Prod.fst).toList @@ -345,7 +346,7 @@ def move_oper_simpCtx : MetaM Simp.Context := do ``min_comm, ``min_assoc, ``min_left_comm -- for `min` ] let simpThms ← simpNames.foldlM (·.addConst ·) ({} : SimpTheorems) - return { simpTheorems := #[simpThms] } + Simp.mkContext {} (simpTheorems := #[simpThms]) /-- `reorderAndSimp mv op instr` takes as input an `MVarId` `mv`, the name `op` of a binary operation and a list of "instructions" `instr` that it passes to `permuteExpr`. diff --git a/Mathlib/Tactic/NormNum/Basic.lean b/Mathlib/Tactic/NormNum/Basic.lean index e5511ceb38673..a1d9e39f9a55b 100644 --- a/Mathlib/Tactic/NormNum/Basic.lean +++ b/Mathlib/Tactic/NormNum/Basic.lean @@ -34,7 +34,7 @@ the unused variable linter can not see usages of variables in set_option linter.unusedVariables false namespace Mathlib -open Lean hiding Rat mkRat +open Lean open Meta namespace Meta.NormNum diff --git a/Mathlib/Tactic/NormNum/BigOperators.lean b/Mathlib/Tactic/NormNum/BigOperators.lean index a5cb43afd77a6..b6851faeaa70e 100644 --- a/Mathlib/Tactic/NormNum/BigOperators.lean +++ b/Mathlib/Tactic/NormNum/BigOperators.lean @@ -38,7 +38,7 @@ In particular, we can't use the plugin on sums containing variables. namespace Mathlib.Meta -open Lean hiding Rat mkRat +open Lean open Meta open Qq diff --git a/Mathlib/Tactic/NormNum/Core.lean b/Mathlib/Tactic/NormNum/Core.lean index 507719e25a60a..65bc7ba6a5a39 100644 --- a/Mathlib/Tactic/NormNum/Core.lean +++ b/Mathlib/Tactic/NormNum/Core.lean @@ -17,7 +17,7 @@ The actual behavior is in `@[norm_num]`-tagged definitions in `Tactic.NormNum.Ba and elsewhere. -/ -open Lean hiding Rat mkRat +open Lean open Lean.Meta Qq Lean.Elab Term /-- Attribute for identifying `norm_num` extensions. -/ @@ -60,9 +60,6 @@ structure NormNums where erased : PHashSet Name := {} deriving Inhabited -/-- Configuration for `DiscrTree`. -/ -def discrTreeConfig : WhnfCoreConfig := {} - /-- Environment extensions for `norm_num` declarations -/ initialize normNumExt : ScopedEnvExtension Entry (Entry × NormNumExt) NormNums ← -- we only need this to deduplicate entries in the DiscrTree @@ -86,7 +83,7 @@ def derive {α : Q(Type u)} (e : Q($α)) (post := false) : MetaM (Result e) := d profileitM Exception "norm_num" (← getOptions) do let s ← saveState let normNums := normNumExt.getState (← getEnv) - let arr ← normNums.tree.getMatch e discrTreeConfig + let arr ← normNums.tree.getMatch e for ext in arr do if (bif post then ext.post else ext.pre) && ! normNums.erased.contains ext.name then try @@ -180,7 +177,7 @@ initialize registerBuiltinAttribute { let e ← elabTerm stx none let (_, _, e) ← lambdaMetaTelescope (← mkLambdaFVars (← getLCtx).getFVars e) return e - DiscrTree.mkPath e discrTreeConfig + DiscrTree.mkPath e normNumExt.add ((keys, declName), ext) kind | _ => throwUnsupportedSyntax erase := fun declName => do @@ -239,7 +236,7 @@ def normNumAt (g : MVarId) (ctx : Simp.Context) (fvarIdsToSimp : Array FVarId) for fvarId in fvarIdsToSimp do let localDecl ← fvarId.getDecl let type ← instantiateMVars localDecl.type - let ctx := { ctx with simpTheorems := ctx.simpTheorems.eraseTheorem (.fvar localDecl.fvarId) } + let ctx := ctx.setSimpTheorems (ctx.simpTheorems.eraseTheorem (.fvar localDecl.fvarId)) let r ← deriveSimp ctx useSimp type match r.proof? with | some _ => @@ -275,13 +272,14 @@ def getSimpContext (cfg args : Syntax) (simpOnly := false) : TacticM Simp.Contex if simpOnly then simpOnlyBuiltins.foldlM (·.addConst ·) {} else getSimpTheorems let mut { ctx, simprocs := _, starArg } ← elabSimpArgs args[0] (eraseLocal := false) (kind := .simp) (simprocs := {}) - { config, simpTheorems := #[simpTheorems], congrTheorems := ← getSimpCongrTheorems } + (← Simp.mkContext config (simpTheorems := #[simpTheorems]) + (congrTheorems := ← getSimpCongrTheorems)) unless starArg do return ctx let mut simpTheorems := ctx.simpTheorems for h in ← getPropHyps do unless simpTheorems.isErased (.fvar h) do simpTheorems ← simpTheorems.addTheorem (.fvar h) (← h.getDecl).toExpr - pure { ctx with simpTheorems } + return ctx.setSimpTheorems simpTheorems open Elab.Tactic in /-- diff --git a/Mathlib/Tactic/NormNum/DivMod.lean b/Mathlib/Tactic/NormNum/DivMod.lean index 17f454c519354..90f6938e16c6b 100644 --- a/Mathlib/Tactic/NormNum/DivMod.lean +++ b/Mathlib/Tactic/NormNum/DivMod.lean @@ -22,7 +22,7 @@ the unused variable linter can not see usages of variables in set_option linter.unusedVariables false namespace Mathlib -open Lean hiding Rat mkRat +open Lean open Meta namespace Meta.NormNum diff --git a/Mathlib/Tactic/NormNum/OfScientific.lean b/Mathlib/Tactic/NormNum/OfScientific.lean index c494615ef34f0..022e8ff81ca0e 100644 --- a/Mathlib/Tactic/NormNum/OfScientific.lean +++ b/Mathlib/Tactic/NormNum/OfScientific.lean @@ -11,7 +11,7 @@ import Mathlib.Data.Rat.Cast.Lemmas -/ namespace Mathlib -open Lean hiding Rat mkRat +open Lean open Meta namespace Meta.NormNum diff --git a/Mathlib/Tactic/NormNum/Pow.lean b/Mathlib/Tactic/NormNum/Pow.lean index 710bb86159091..467a97ad68b3f 100644 --- a/Mathlib/Tactic/NormNum/Pow.lean +++ b/Mathlib/Tactic/NormNum/Pow.lean @@ -19,7 +19,7 @@ the unused variable linter can not see usages of variables in set_option linter.unusedVariables false namespace Mathlib -open Lean hiding Rat mkRat +open Lean open Meta namespace Meta.NormNum diff --git a/Mathlib/Tactic/NormNum/PowMod.lean b/Mathlib/Tactic/NormNum/PowMod.lean index f50964e19a2b3..b3ac0901c7829 100644 --- a/Mathlib/Tactic/NormNum/PowMod.lean +++ b/Mathlib/Tactic/NormNum/PowMod.lean @@ -24,7 +24,7 @@ The approach taken here is identical to (and copied from) the development in `No set_option autoImplicit true namespace Mathlib -open Lean hiding Rat mkRat +open Lean open Meta namespace Meta.NormNum diff --git a/Mathlib/Tactic/NormNum/Result.lean b/Mathlib/Tactic/NormNum/Result.lean index 85a73f0db9eca..aa37896d595c5 100644 --- a/Mathlib/Tactic/NormNum/Result.lean +++ b/Mathlib/Tactic/NormNum/Result.lean @@ -26,7 +26,7 @@ or is either `true` or `false`. universe u variable {α : Type u} -open Lean hiding Rat mkRat +open Lean open Lean.Meta Qq Lean.Elab Term namespace Mathlib diff --git a/Mathlib/Tactic/Polyrith.lean b/Mathlib/Tactic/Polyrith.lean index 50e42c8eecbd1..91b368722375e 100644 --- a/Mathlib/Tactic/Polyrith.lean +++ b/Mathlib/Tactic/Polyrith.lean @@ -60,7 +60,7 @@ remember to force recompilation of any files that call `polyrith`. -/ namespace Mathlib.Tactic.Polyrith -open Lean hiding Rat +open Lean open Meta Ring Qq PrettyPrinter AtomM initialize registerTraceClass `Meta.Tactic.polyrith diff --git a/Mathlib/Tactic/Positivity/Core.lean b/Mathlib/Tactic/Positivity/Core.lean index 0b6a352c4ae01..a56d312db10c5 100644 --- a/Mathlib/Tactic/Positivity/Core.lean +++ b/Mathlib/Tactic/Positivity/Core.lean @@ -20,7 +20,7 @@ The actual behavior is in `@[positivity]`-tagged definitions in `Tactic.Positivi and elsewhere. -/ -open Lean hiding Rat +open Lean open Lean.Meta Qq Lean.Elab Term /-- Attribute for identifying `positivity` extensions. -/ @@ -71,9 +71,6 @@ def mkPositivityExt (n : Name) : ImportM PositivityExt := do let { env, opts, .. } ← read IO.ofExcept <| unsafe env.evalConstCheck PositivityExt opts ``PositivityExt n -/-- Configuration for `DiscrTree`. -/ -def discrTreeConfig : WhnfCoreConfig := {} - /-- Each `positivity` extension is labelled with a collection of patterns which determine the expressions to which it should be applied. -/ abbrev Entry := Array (Array DiscrTree.Key) × Name @@ -113,7 +110,7 @@ initialize registerBuiltinAttribute { let e ← elabTerm stx none let (_, _, e) ← lambdaMetaTelescope (← mkLambdaFVars (← getLCtx).getFVars e) return e - DiscrTree.mkPath e discrTreeConfig + DiscrTree.mkPath e setEnv <| positivityExt.addEntry env ((keys, declName), ext) | _ => throwUnsupportedSyntax } @@ -318,7 +315,7 @@ def orElse {e : Q($α)} (t₁ : Strictness zα pα e) (t₂ : MetaM (Strictness def core (e : Q($α)) : MetaM (Strictness zα pα e) := do let mut result := .none trace[Tactic.positivity] "trying to prove positivity of {e}" - for ext in ← (positivityExt.getState (← getEnv)).2.getMatch e discrTreeConfig do + for ext in ← (positivityExt.getState (← getEnv)).2.getMatch e do try result ← orElse result <| ext.eval zα pα e catch err => diff --git a/Mathlib/Tactic/Propose.lean b/Mathlib/Tactic/Propose.lean index 4169210c6c5ce..debb88c19ec1b 100644 --- a/Mathlib/Tactic/Propose.lean +++ b/Mathlib/Tactic/Propose.lean @@ -39,9 +39,6 @@ open Lean Meta Batteries.Tactic Tactic.TryThis initialize registerTraceClass `Tactic.propose -/-- Configuration for `DiscrTree`. -/ -def discrTreeConfig : WhnfCoreConfig := {} - initialize proposeLemmas : DeclCache (DiscrTree Name) ← DeclCache.mk "have?: init cache" failure {} fun name constInfo lemmas => do if constInfo.isUnsafe then return lemmas @@ -50,7 +47,7 @@ initialize proposeLemmas : DeclCache (DiscrTree Name) ← let (mvars, _, _) ← forallMetaTelescope constInfo.type let mut lemmas := lemmas for m in mvars do - lemmas ← lemmas.insertIfSpecific (← inferType m) name discrTreeConfig + lemmas ← lemmas.insertIfSpecific (← inferType m) name pure lemmas open Lean.Meta.SolveByElim in @@ -80,7 +77,7 @@ def propose (lemmas : DiscrTree Name) (type : Expr) (required : Array Expr) (solveByElimDepth := 15) : MetaM (Array (Name × Expr)) := do guard !required.isEmpty let ty ← whnfR (← instantiateMVars (← inferType required[0]!)) - let candidates ← lemmas.getMatch ty discrTreeConfig + let candidates ← lemmas.getMatch ty candidates.filterMapM fun lem : Name => try trace[Tactic.propose] "considering {lem}" diff --git a/Mathlib/Tactic/PushNeg.lean b/Mathlib/Tactic/PushNeg.lean index 39b7ff1d2cd6b..4c59bb174597f 100644 --- a/Mathlib/Tactic/PushNeg.lean +++ b/Mathlib/Tactic/PushNeg.lean @@ -144,10 +144,9 @@ partial def transformNegation (e : Expr) : SimpM Simp.Step := do /-- Common entry point to `push_neg` as a conv. -/ def pushNegCore (tgt : Expr) : MetaM Simp.Result := do - let myctx : Simp.Context := - { config := { eta := true, zeta := false, proj := false }, - simpTheorems := #[ ] - congrTheorems := (← getSimpCongrTheorems) } + let myctx : Simp.Context ← Simp.mkContext { eta := true, zeta := false, proj := false } + (simpTheorems := #[ ]) + (congrTheorems := (← getSimpCongrTheorems)) (·.1) <$> Simp.main tgt myctx (methods := { pre := transformNegation }) /-- diff --git a/Mathlib/Tactic/ReduceModChar.lean b/Mathlib/Tactic/ReduceModChar.lean index a7689ecfbbe80..ed98906d7949f 100644 --- a/Mathlib/Tactic/ReduceModChar.lean +++ b/Mathlib/Tactic/ReduceModChar.lean @@ -251,11 +251,8 @@ partial def derive (expensive := false) (e : Expr) : MetaM Simp.Result := do let ext ← match ext? with | some ext => pure ext | none => throwError "internal error: reduce_mod_char not registered as simp extension" - let ctx : Simp.Context := { - config := config, - congrTheorems := congrTheorems, - simpTheorems := #[← ext.getTheorems] - } + let ctx ← Simp.mkContext config (congrTheorems := congrTheorems) + (simpTheorems := #[← ext.getTheorems]) let discharge := Mathlib.Meta.NormNum.discharge ctx let r : Simp.Result := {expr := e} let pre := Simp.preDefault #[] >> fun e => diff --git a/Mathlib/Tactic/Relation/Rfl.lean b/Mathlib/Tactic/Relation/Rfl.lean index b35a56bd7c5c9..2abf31b41b4df 100644 --- a/Mathlib/Tactic/Relation/Rfl.lean +++ b/Mathlib/Tactic/Relation/Rfl.lean @@ -36,7 +36,7 @@ def _root_.Lean.Expr.relSidesIfRefl? (e : Expr) : MetaM (Option (Name × Expr × if let some (_, lhs, _, rhs) := e.heq? then return (``HEq, lhs, rhs) if let .app (.app rel lhs) rhs := e then - unless (← (reflExt.getState (← getEnv)).getMatch rel reflExt.config).isEmpty do + unless (← (reflExt.getState (← getEnv)).getMatch rel).isEmpty do match rel.getAppFn.constName? with | some n => return some (n, lhs, rhs) | none => return none diff --git a/Mathlib/Tactic/Relation/Symm.lean b/Mathlib/Tactic/Relation/Symm.lean index 8d33854268714..aa240b55b3d28 100644 --- a/Mathlib/Tactic/Relation/Symm.lean +++ b/Mathlib/Tactic/Relation/Symm.lean @@ -27,7 +27,7 @@ def _root_.Lean.Expr.relSidesIfSymm? (e : Expr) : MetaM (Option (Name × Expr × if let some (_, lhs, _, rhs) := e.heq? then return (``HEq, lhs, rhs) if let .app (.app rel lhs) rhs := e then - unless (← (symmExt.getState (← getEnv)).getMatch rel symmExt.config).isEmpty do + unless (← (symmExt.getState (← getEnv)).getMatch rel).isEmpty do match rel.getAppFn.constName? with | some n => return some (n, lhs, rhs) | none => return none diff --git a/Mathlib/Tactic/Ring/RingNF.lean b/Mathlib/Tactic/Ring/RingNF.lean index 6b9c13c7cca7a..3d57d68676c46 100644 --- a/Mathlib/Tactic/Ring/RingNF.lean +++ b/Mathlib/Tactic/Ring/RingNF.lean @@ -19,7 +19,7 @@ such as `sin (x + y) + sin (y + x) = 2 * sin (x + y)`. -/ namespace Mathlib.Tactic -open Lean hiding Rat +open Lean open Qq Meta namespace Ring @@ -133,10 +133,9 @@ Runs a tactic in the `RingNF.M` monad, given initial data: -/ partial def M.run {α : Type} (s : IO.Ref AtomM.State) (cfg : RingNF.Config) (x : M α) : MetaM α := do - let ctx := { - simpTheorems := #[← Elab.Tactic.simpOnlyBuiltins.foldlM (·.addConst ·) {}] - congrTheorems := ← getSimpCongrTheorems - config.singlePass := cfg.mode matches .raw } + let ctx ← Simp.mkContext { singlePass := cfg.mode matches .raw } + (simpTheorems := #[← Elab.Tactic.simpOnlyBuiltins.foldlM (·.addConst ·) {}]) + (congrTheorems := ← getSimpCongrTheorems) let simp ← match cfg.mode with | .raw => pure pure | .SOP => @@ -145,7 +144,7 @@ partial def M.run ``_root_.pow_one, ``mul_neg, ``add_neg].foldlM (·.addConst ·) thms let thms ← [``nat_rawCast_0, ``nat_rawCast_1, ``nat_rawCast_2, ``int_rawCast_neg, ``rat_rawCast_neg, ``rat_rawCast_pos].foldlM (·.addConst · (post := false)) thms - let ctx' := { ctx with simpTheorems := #[thms] } + let ctx' := ctx.setSimpTheorems #[thms] pure fun r' : Simp.Result ↦ do r'.mkEqTrans (← Simp.main r'.expr ctx' (methods := Lean.Meta.Simp.mkDefaultMethodsCore {})).1 let nctx := { ctx, simp } diff --git a/Mathlib/Tactic/SimpIntro.lean b/Mathlib/Tactic/SimpIntro.lean index f2f3ff6f8b497..a816ba1df65b2 100644 --- a/Mathlib/Tactic/SimpIntro.lean +++ b/Mathlib/Tactic/SimpIntro.lean @@ -31,7 +31,7 @@ partial def simpIntroCore (g : MVarId) (ctx : Simp.Context) (simprocs : Simp.Sim let withFVar := fun (fvar, g) ↦ g.withContext do Term.addLocalVarInfo var (mkFVar fvar) let simpTheorems ← ctx.simpTheorems.addTheorem (.fvar fvar) (.fvar fvar) - simpIntroCore g { ctx with simpTheorems } simprocs discharge? more ids' + simpIntroCore g (ctx.setSimpTheorems simpTheorems) simprocs discharge? more ids' match t with | .letE .. => withFVar (← g.intro n) | .forallE (body := body) .. => diff --git a/Mathlib/Tactic/Simps/Basic.lean b/Mathlib/Tactic/Simps/Basic.lean index dac4370d3235c..9bff6a59a5119 100644 --- a/Mathlib/Tactic/Simps/Basic.lean +++ b/Mathlib/Tactic/Simps/Basic.lean @@ -82,10 +82,9 @@ def mkSimpContextResult (cfg : Meta.Simp.Config := {}) (simpOnly := false) (kind getSimpTheorems let simprocs := #[← if simpOnly then pure {} else Simp.getSimprocs] let congrTheorems ← getSimpCongrTheorems - let ctx : Simp.Context := { - config := cfg - simpTheorems := #[simpTheorems], congrTheorems - } + let ctx : Simp.Context ← Simp.mkContext cfg + (simpTheorems := #[simpTheorems]) + (congrTheorems := congrTheorems) if !hasStar then return { ctx, simprocs, dischargeWrapper } else @@ -94,7 +93,7 @@ def mkSimpContextResult (cfg : Meta.Simp.Config := {}) (simpOnly := false) (kind for h in hs do unless simpTheorems.isErased (.fvar h) do simpTheorems ← simpTheorems.addTheorem (.fvar h) (← h.getDecl).toExpr - let ctx := { ctx with simpTheorems } + let ctx := ctx.setSimpTheorems simpTheorems return { ctx, simprocs, dischargeWrapper } /-- Make `Simp.Context` giving data instead of Syntax. Doesn't support arguments. diff --git a/Mathlib/Tactic/SplitIfs.lean b/Mathlib/Tactic/SplitIfs.lean index 49287e1c5ec75..64e0fa989bcac 100644 --- a/Mathlib/Tactic/SplitIfs.lean +++ b/Mathlib/Tactic/SplitIfs.lean @@ -75,7 +75,7 @@ private def discharge? (e : Expr) : SimpM (Option Expr) := do -/ private def reduceIfsAt (loc : Location) : TacticM Unit := do let ctx ← SplitIf.getSimpContext - let ctx := { ctx with config := { ctx.config with failIfUnchanged := false } } + let ctx := ctx.setFailIfUnchanged false let _ ← simpLocation ctx (← ({} : Simp.SimprocsArray).add `reduceCtorEq false) discharge? loc pure () diff --git a/Mathlib/Tactic/Spread.lean b/Mathlib/Tactic/Spread.lean index 0252a2fa396db..36c7fb615c6b7 100644 --- a/Mathlib/Tactic/Spread.lean +++ b/Mathlib/Tactic/Spread.lean @@ -66,8 +66,6 @@ macro_rules spreads := spreads.push arg else newFields := newFields.push field - | `(structInstFieldAbbrev| $_:ident) => - newFields := newFields.push field | _ => throwUnsupported diff --git a/Mathlib/Tactic/ToAdditive/Frontend.lean b/Mathlib/Tactic/ToAdditive/Frontend.lean index 25e3af975e0bf..b0278e30a9e9d 100644 --- a/Mathlib/Tactic/ToAdditive/Frontend.lean +++ b/Mathlib/Tactic/ToAdditive/Frontend.lean @@ -866,7 +866,7 @@ def additivizeLemmas {m : Type → Type} [Monad m] [MonadError m] [MonadLiftT Co for (nm, lemmas) in names.zip auxLemmas do unless lemmas.size == nLemmas do throwError "{names[0]!} and {nm} do not generate the same number of {desc}." - for (srcLemmas, tgtLemmas) in auxLemmas.zip <| auxLemmas.eraseIdx 0 do + for (srcLemmas, tgtLemmas) in auxLemmas.zip <| auxLemmas.eraseIdx! 0 do for (srcLemma, tgtLemma) in srcLemmas.zip tgtLemmas do insertTranslation srcLemma tgtLemma diff --git a/Mathlib/Tactic/Variable.lean b/Mathlib/Tactic/Variable.lean index c506333c3c075..3f885e361c071 100644 --- a/Mathlib/Tactic/Variable.lean +++ b/Mathlib/Tactic/Variable.lean @@ -159,7 +159,7 @@ partial def completeBinders' (maxSteps : Nat) (gas : Nat) (binders : TSyntaxArray ``bracketedBinder) (toOmit : Array Bool) (i : Nat) : TermElabM (TSyntaxArray ``bracketedBinder × Array Bool) := do - if 0 < gas && i < binders.size then + if h : 0 < gas ∧ i < binders.size then let binder := binders[i]! trace[«variable?»] "\ Have {(← getLCtx).getFVarIds.size} fvars and {(← getLocalInstances).size} local instances. \ @@ -176,7 +176,7 @@ partial def completeBinders' (maxSteps : Nat) (gas : Nat) in binders since they are generated by pretty printing unsatisfied dependencies.\n\n\ Current variable command:{indentD (← `(command| variable $binders'*))}\n\n\ Local context for the unsatisfied dependency:{goalMsg}" - let binders := binders.insertAt! i binder' + let binders := binders.insertIdx i binder' completeBinders' maxSteps (gas - 1) checkRedundant binders toOmit i else let lctx ← getLCtx diff --git a/Mathlib/Tactic/WLOG.lean b/Mathlib/Tactic/WLOG.lean index 461949d04ab7b..cad9a19bf5bfd 100644 --- a/Mathlib/Tactic/WLOG.lean +++ b/Mathlib/Tactic/WLOG.lean @@ -77,7 +77,8 @@ def _root_.Lean.MVarId.wlog (goal : MVarId) (h : Option Name) (P : Expr) let (revertedFVars, HType) ← liftMkBindingM fun ctx => (do let f ← collectForwardDeps lctx fvars let revertedFVars := filterOutImplementationDetails lctx (f.map Expr.fvarId!) - let HType ← withFreshCache do mkAuxMVarType lctx (revertedFVars.map Expr.fvar) .natural HSuffix + let HType ← withFreshCache do + mkAuxMVarType lctx (revertedFVars.map Expr.fvar) .natural HSuffix (usedLetOnly := true) return (revertedFVars, HType)) { preserveOrder := false, mainModule := ctx.mainModule } /- Set up the goal which will suppose `h`; this begins as a goal with type H (hence HExpr), and h diff --git a/Mathlib/Testing/Plausible/Functions.lean b/Mathlib/Testing/Plausible/Functions.lean index a9bdf24d8ccb4..ba4209766edf6 100644 --- a/Mathlib/Testing/Plausible/Functions.lean +++ b/Mathlib/Testing/Plausible/Functions.lean @@ -202,7 +202,7 @@ theorem List.applyId_zip_eq [DecidableEq α] {xs ys : List α} (h₀ : List.Nodu simp only [getElem?_cons_succ, zip_cons_cons, applyId_cons] at h₂ ⊢ rw [if_neg] · apply xs_ih <;> solve_by_elim [Nat.succ.inj] - · apply h₀; apply List.getElem?_mem h₂ + · apply h₀; apply List.mem_of_getElem? h₂ theorem applyId_mem_iff [DecidableEq α] {xs ys : List α} (h₀ : List.Nodup xs) (h₁ : xs ~ ys) (x : α) : List.applyId.{u} (xs.zip ys) x ∈ ys ↔ x ∈ xs := by diff --git a/Mathlib/Topology/Algebra/PontryaginDual.lean b/Mathlib/Topology/Algebra/PontryaginDual.lean index 4229049cf7321..2b24a7b06909d 100644 --- a/Mathlib/Topology/Algebra/PontryaginDual.lean +++ b/Mathlib/Topology/Algebra/PontryaginDual.lean @@ -73,7 +73,7 @@ instance [LocallyCompactSpace H] : LocallyCompactSpace (PontryaginDual H) := by refine lt_of_lt_of_le ht ?_ rw [div_le_iff₀' (pow_pos two_pos _), ← div_le_iff₀ hx] refine (Nat.le_ceil (Real.pi / x)).trans ?_ - exact_mod_cast (Nat.le_succ _).trans (Nat.lt_two_pow _).le + exact_mod_cast (Nat.le_succ _).trans Nat.lt_two_pow_self.le variable {A B C G} diff --git a/Mathlib/Topology/Gluing.lean b/Mathlib/Topology/Gluing.lean index c5385f98801ee..6c4f75721330d 100644 --- a/Mathlib/Topology/Gluing.lean +++ b/Mathlib/Topology/Gluing.lean @@ -83,7 +83,7 @@ Most of the times it would be easier to use the constructor `TopCat.GlueData.mk' conditions are stated in a less categorical way. -/ -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] -structure GlueData extends GlueData TopCat where +structure GlueData extends CategoryTheory.GlueData TopCat where f_open : ∀ i j, IsOpenEmbedding (f i j) f_mono i j := (TopCat.mono_iff_injective _).mpr (f_open i j).isEmbedding.injective diff --git a/Mathlib/Topology/List.lean b/Mathlib/Topology/List.lean index 0620deab85315..4ca0d9581282a 100644 --- a/Mathlib/Topology/List.lean +++ b/Mathlib/Topology/List.lean @@ -184,16 +184,16 @@ open List open Mathlib (Vector) -instance (n : ℕ) : TopologicalSpace (Vector α n) := by unfold Vector; infer_instance +instance (n : ℕ) : TopologicalSpace (Mathlib.Vector α n) := by unfold Mathlib.Vector; infer_instance -theorem tendsto_cons {n : ℕ} {a : α} {l : Vector α n} : - Tendsto (fun p : α × Vector α n => p.1 ::ᵥ p.2) (𝓝 a ×ˢ 𝓝 l) (𝓝 (a ::ᵥ l)) := by +theorem tendsto_cons {n : ℕ} {a : α} {l : Mathlib.Vector α n} : + Tendsto (fun p : α × Mathlib.Vector α n => p.1 ::ᵥ p.2) (𝓝 a ×ˢ 𝓝 l) (𝓝 (a ::ᵥ l)) := by rw [tendsto_subtype_rng, Vector.cons_val] exact tendsto_fst.cons (Tendsto.comp continuousAt_subtype_val tendsto_snd) theorem tendsto_insertIdx {n : ℕ} {i : Fin (n + 1)} {a : α} : - ∀ {l : Vector α n}, - Tendsto (fun p : α × Vector α n => Vector.insertIdx p.1 i p.2) (𝓝 a ×ˢ 𝓝 l) + ∀ {l : Mathlib.Vector α n}, + Tendsto (fun p : α × Mathlib.Vector α n => Vector.insertIdx p.1 i p.2) (𝓝 a ×ˢ 𝓝 l) (𝓝 (Vector.insertIdx a i l)) | ⟨l, hl⟩ => by rw [Vector.insertIdx, tendsto_subtype_rng] @@ -204,29 +204,29 @@ theorem tendsto_insertIdx {n : ℕ} {i : Fin (n + 1)} {a : α} : /-- Continuity of `Vector.insertIdx`. -/ theorem continuous_insertIdx' {n : ℕ} {i : Fin (n + 1)} : - Continuous fun p : α × Vector α n => Vector.insertIdx p.1 i p.2 := + Continuous fun p : α × Mathlib.Vector α n => Vector.insertIdx p.1 i p.2 := continuous_iff_continuousAt.mpr fun ⟨a, l⟩ => by rw [ContinuousAt, nhds_prod_eq]; exact tendsto_insertIdx @[deprecated (since := "2024-10-21")] alias continuous_insertNth' := continuous_insertIdx' -theorem continuous_insertIdx {n : ℕ} {i : Fin (n + 1)} {f : β → α} {g : β → Vector α n} +theorem continuous_insertIdx {n : ℕ} {i : Fin (n + 1)} {f : β → α} {g : β → Mathlib.Vector α n} (hf : Continuous f) (hg : Continuous g) : Continuous fun b => Vector.insertIdx (f b) i (g b) := continuous_insertIdx'.comp (hf.prod_mk hg : _) @[deprecated (since := "2024-10-21")] alias continuous_insertNth := continuous_insertIdx theorem continuousAt_eraseIdx {n : ℕ} {i : Fin (n + 1)} : - ∀ {l : Vector α (n + 1)}, ContinuousAt (Vector.eraseIdx i) l + ∀ {l : Mathlib.Vector α (n + 1)}, ContinuousAt (Mathlib.Vector.eraseIdx i) l | ⟨l, hl⟩ => by - rw [ContinuousAt, Vector.eraseIdx, tendsto_subtype_rng] + rw [ContinuousAt, Mathlib.Vector.eraseIdx, tendsto_subtype_rng] simp only [Vector.eraseIdx_val] exact Tendsto.comp List.tendsto_eraseIdx continuousAt_subtype_val @[deprecated (since := "2024-05-04")] alias continuousAt_removeNth := continuousAt_eraseIdx theorem continuous_eraseIdx {n : ℕ} {i : Fin (n + 1)} : - Continuous (Vector.eraseIdx i : Vector α (n + 1) → Vector α n) := + Continuous (Mathlib.Vector.eraseIdx i : Mathlib.Vector α (n + 1) → Mathlib.Vector α n) := continuous_iff_continuousAt.mpr fun ⟨_a, _l⟩ => continuousAt_eraseIdx @[deprecated (since := "2024-05-04")] alias continuous_removeNth := continuous_eraseIdx diff --git a/Mathlib/Topology/OmegaCompletePartialOrder.lean b/Mathlib/Topology/OmegaCompletePartialOrder.lean index c7890d3939837..cce0eb4e71f4a 100644 --- a/Mathlib/Topology/OmegaCompletePartialOrder.lean +++ b/Mathlib/Topology/OmegaCompletePartialOrder.lean @@ -106,7 +106,7 @@ theorem notBelow_isOpen : IsOpen (notBelow y) := by end notBelow -open Scott hiding IsOpen +open Scott hiding IsOpen IsOpen.isUpperSet open OmegaCompletePartialOrder diff --git a/Mathlib/Util/DischargerAsTactic.lean b/Mathlib/Util/DischargerAsTactic.lean index b7662072b5f0a..cbac978463f1d 100644 --- a/Mathlib/Util/DischargerAsTactic.lean +++ b/Mathlib/Util/DischargerAsTactic.lean @@ -22,7 +22,7 @@ so that it can be passed as an argument to `simp (discharger := foo)`. This is inverse to `mkDischargeWrapper`. -/ def wrapSimpDischarger (dis : Simp.Discharge) : TacticM Unit := do let eS : Lean.Meta.Simp.State := {} - let eC : Lean.Meta.Simp.Context := {} + let eC : Lean.Meta.Simp.Context := ← Simp.mkContext {} let eM : Lean.Meta.Simp.Methods := {} let (some a, _) ← liftM <| StateRefT'.run (ReaderT.run (ReaderT.run (dis <| ← getMainTarget) eM.toMethodsRef) eC) eS | failure diff --git a/MathlibTest/Expr.lean b/MathlibTest/Expr.lean index 1e4b9fd11388d..84efba4ecb2f9 100644 --- a/MathlibTest/Expr.lean +++ b/MathlibTest/Expr.lean @@ -13,7 +13,7 @@ def reorderLastArguments : Expr → Expr := Expr.replaceRec fun r e ↦ let n := e.getAppNumArgs if n ≥ 2 then - mkAppN e.getAppFn <| e.getAppArgs.map r |>.swap! (n - 1) (n - 2) + mkAppN e.getAppFn <| e.getAppArgs.map r |>.swapIfInBounds (n - 1) (n - 2) else none diff --git a/MathlibTest/NthRewrite.lean b/MathlibTest/NthRewrite.lean index 287039007dce6..3b82ffdb5998a 100644 --- a/MathlibTest/NthRewrite.lean +++ b/MathlibTest/NthRewrite.lean @@ -16,7 +16,7 @@ example [AddZeroClass G] {a : G} : a + a = a + (a + 0) := by structure F where (a : ℕ) - (v : Vector ℕ a) + (v : Mathlib.Vector ℕ a) (p : v.val = []) example (f : F) : f.v.val = [] := by diff --git a/MathlibTest/Simps.lean b/MathlibTest/Simps.lean index 8def1aa96d6ef..2c3e0d2f69483 100644 --- a/MathlibTest/Simps.lean +++ b/MathlibTest/Simps.lean @@ -1203,7 +1203,7 @@ structure Foo where @[simps field field_2 field_9_fst] def myFoo : Foo := ⟨1, ⟨1, 1⟩, 1⟩ -structure Prod (X Y : Type _) extends Prod X Y +structure Prod (X Y : Type _) extends _root_.Prod X Y structure Prod2 (X Y : Type _) extends Prod X Y diff --git a/MathlibTest/cc.lean b/MathlibTest/cc.lean index 4d0f3eb630b40..26e0d6ac83f4d 100644 --- a/MathlibTest/cc.lean +++ b/MathlibTest/cc.lean @@ -220,17 +220,21 @@ universe u open Mathlib -axiom app : {α : Type u} → {n m : Nat} → Vector α m → Vector α n → Vector α (m + n) +axiom app : {α : Type u} → {n m : Nat} → + Mathlib.Vector α m → Mathlib.Vector α n → Mathlib.Vector α (m + n) -example (n1 n2 n3 : Nat) (v1 w1 : Vector Nat n1) (w1' : Vector Nat n3) (v2 w2 : Vector Nat n2) : +example (n1 n2 n3 : Nat) + (v1 w1 : Mathlib.Vector Nat n1) (w1' : Mathlib.Vector Nat n3) (v2 w2 : Mathlib.Vector Nat n2) : n1 = n3 → v1 = w1 → HEq w1 w1' → v2 = w2 → HEq (app v1 v2) (app w1' w2) := by cc -example (n1 n2 n3 : Nat) (v1 w1 : Vector Nat n1) (w1' : Vector Nat n3) (v2 w2 : Vector Nat n2) : +example (n1 n2 n3 : Nat) + (v1 w1 : Mathlib.Vector Nat n1) (w1' : Mathlib.Vector Nat n3) (v2 w2 : Mathlib.Vector Nat n2) : HEq n1 n3 → v1 = w1 → HEq w1 w1' → HEq v2 w2 → HEq (app v1 v2) (app w1' w2) := by cc -example (n1 n2 n3 : Nat) (v1 w1 v : Vector Nat n1) (w1' : Vector Nat n3) (v2 w2 w : Vector Nat n2) : +example (n1 n2 n3 : Nat) + (v1 w1 v : Mathlib.Vector Nat n1) (w1' : Mathlib.Vector Nat n3) (v2 w2 w : Mathlib.Vector Nat n2) : HEq n1 n3 → v1 = w1 → HEq w1 w1' → HEq v2 w2 → HEq (app w1' w2) (app v w) → app v1 v2 = app v w := by cc diff --git a/lake-manifest.json b/lake-manifest.json index 81bcd7da5e4e8..0f235638e282f 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,10 +5,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "42dc02bdbc5d0c2f395718462a76c3d87318f7fa", + "rev": "8e5cb8d424df462f84997dd68af6f40e347c3e35", "name": "plausible", "manifestFile": "lake-manifest.json", - "inputRev": "main", + "inputRev": "v4.15.0-rc1", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover-community/LeanSearchClient", @@ -25,47 +25,47 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "519e509a28864af5bed98033dd33b95cf08e9aa7", + "rev": "ed3b856bd8893ade75cafe13e8544d4c2660f377", "name": "importGraph", "manifestFile": "lake-manifest.json", - "inputRev": "v4.14.0", + "inputRev": "v4.15.0-rc1", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover-community/ProofWidgets4", "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "68280daef58803f68368eb2e53046dabcd270c9d", + "rev": "2b000e02d50394af68cfb4770a291113d94801b5", "name": "proofwidgets", "manifestFile": "lake-manifest.json", - "inputRev": "v0.0.47", + "inputRev": "v0.0.48", "inherited": false, "configFile": "lakefile.lean"}, {"url": "https://github.com/leanprover-community/aesop", "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "5a0ec8588855265ade536f35bcdcf0fb24fd6030", + "rev": "8b6048aa0a4a4b6bcf83597802d8dee734e64b7e", "name": "aesop", "manifestFile": "lake-manifest.json", - "inputRev": "v4.14.0", + "inputRev": "v4.15.0-rc1", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover-community/quote4", "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "303b23fbcea94ac4f96e590c1cad6618fd4f5f41", + "rev": "ad942fdf0b15c38bface6acbb01d63855a2519ac", "name": "Qq", "manifestFile": "lake-manifest.json", - "inputRev": "master", + "inputRev": "v4.14.0", "inherited": false, "configFile": "lakefile.lean"}, {"url": "https://github.com/leanprover-community/batteries", "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "8d6c853f11a5172efa0e96b9f2be1a83d861cdd9", + "rev": "7805acf1864ba1a2e359f86a8f092ccf1438ad83", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "v4.14.0", @@ -75,7 +75,7 @@ "type": "git", "subDir": null, "scope": "leanprover", - "rev": "726b3c9ad13acca724d4651f14afc4804a7b0e4d", + "rev": "0c8ea32a15a4f74143e4e1e107ba2c412adb90fd", "name": "Cli", "manifestFile": "lake-manifest.json", "inputRev": "main", diff --git a/lakefile.lean b/lakefile.lean index 4f30670824859..27489fbd32314 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -7,14 +7,13 @@ open Lake DSL ## Mathlib dependencies on upstream projects -/ -require "leanprover-community" / "batteries" @ git "v4.14.0" -require "leanprover-community" / "Qq" @ git "master" -require "leanprover-community" / "aesop" @ git "v4.14.0" -require "leanprover-community" / "proofwidgets" @ git "v0.0.47" -require "leanprover-community" / "importGraph" @ git "v4.14.0" +require "leanprover-community" / "batteries" @ git "main" +require "leanprover-community" / "Qq" @ git "v4.14.0" +require "leanprover-community" / "aesop" @ git "v4.15.0-rc1" +require "leanprover-community" / "proofwidgets" @ git "v0.0.48" +require "leanprover-community" / "importGraph" @ git "v4.15.0-rc1" require "leanprover-community" / "LeanSearchClient" @ git "main" - from git "https://github.com/leanprover-community/LeanSearchClient" @ "main" -require "leanprover-community" / "plausible" @ git "main" +require "leanprover-community" / "plausible" @ git "v4.15.0-rc1" /-! ## Options for building mathlib diff --git a/lean-toolchain b/lean-toolchain index 401bc146f623c..cf25a9816f061 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.14.0 \ No newline at end of file +leanprover/lean4:v4.15.0-rc1 From f23ba2e3090396e5752091990b14f5f045f7305e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 2 Dec 2024 05:38:34 +0000 Subject: [PATCH 457/829] feat(GroupTheory): some lemmas about blocks (#19399) From Kneser (LeanCamCombi) --- Mathlib/GroupTheory/GroupAction/Blocks.lean | 36 ++++++++++++++++++++- 1 file changed, 35 insertions(+), 1 deletion(-) diff --git a/Mathlib/GroupTheory/GroupAction/Blocks.lean b/Mathlib/GroupTheory/GroupAction/Blocks.lean index 39dc66ecbf32e..99a9a8640fb29 100644 --- a/Mathlib/GroupTheory/GroupAction/Blocks.lean +++ b/Mathlib/GroupTheory/GroupAction/Blocks.lean @@ -106,7 +106,7 @@ def IsTrivialBlock (B : Set X) := B.Subsingleton ∨ B = univ "A set `B` is a `G`-block iff the sets of the form `g +ᵥ B` are pairwise equal or disjoint. "] def IsBlock (B : Set X) := ∀ ⦃g₁ g₂ : G⦄, g₁ • B ≠ g₂ • B → Disjoint (g₁ • B) (g₂ • B) -variable {G} +variable {G} {s : Set G} {g g₁ g₂ : G} @[to_additive] lemma isBlock_iff_smul_eq_smul_of_nonempty : @@ -122,6 +122,27 @@ lemma isBlock_iff_smul_eq_smul_or_disjoint : IsBlock G B ↔ ∀ g₁ g₂ : G, g₁ • B = g₂ • B ∨ Disjoint (g₁ • B) (g₂ • B) := forall₂_congr fun _ _ ↦ or_iff_not_imp_left.symm +@[to_additive] +lemma IsBlock.smul_eq_smul_of_subset (hB : IsBlock G B) (hg : g₁ • B ⊆ g₂ • B) : + g₁ • B = g₂ • B := by + by_contra! hg' + obtain rfl : B = ∅ := by simpa using (hB hg').eq_bot_of_le hg + simp at hg' + +@[to_additive] +lemma IsBlock.not_smul_set_ssubset_smul_set (hB : IsBlock G B) : ¬ g₁ • B ⊂ g₂ • B := + fun hab ↦ hab.ne <| hB.smul_eq_smul_of_subset hab.subset + +@[to_additive] +lemma IsBlock.disjoint_smul_set_smul (hB : IsBlock G B) (hgs : ¬ g • B ⊆ s • B) : + Disjoint (g • B) (s • B) := by + rw [← iUnion_smul_set, disjoint_iUnion₂_right] + exact fun b hb ↦ hB fun h ↦ hgs <| h.trans_subset <| smul_set_subset_smul hb + +@[to_additive] +lemma IsBlock.disjoint_smul_smul_set (hB : IsBlock G B) (hgs : ¬ g • B ⊆ s • B) : + Disjoint (s • B) (g • B) := (hB.disjoint_smul_set_smul hgs).symm + alias ⟨IsBlock.smul_eq_smul_of_nonempty, _⟩ := isBlock_iff_smul_eq_smul_of_nonempty alias ⟨IsBlock.pairwiseDisjoint_range_smul, _⟩ := isBlock_iff_pairwiseDisjoint_range_smul alias ⟨IsBlock.smul_eq_smul_or_disjoint, _⟩ := isBlock_iff_smul_eq_smul_or_disjoint @@ -150,6 +171,19 @@ lemma IsFixedBlock.isInvariantBlock (hB : IsFixedBlock G B) : IsInvariantBlock G end SMul +section Monoid +variable {M X : Type*} [Monoid M] [MulAction M X] {B : Set X} {s : Set M} + +@[to_additive] +lemma IsBlock.disjoint_smul_right (hB : IsBlock M B) (hs : ¬ B ⊆ s • B) : Disjoint B (s • B) := by + simpa using hB.disjoint_smul_set_smul (g := 1) (by simpa using hs) + +@[to_additive] +lemma IsBlock.disjoint_smul_left (hB : IsBlock M B) (hs : ¬ B ⊆ s • B) : Disjoint (s • B) B := + (hB.disjoint_smul_right hs).symm + +end Monoid + section Group variable {G : Type*} [Group G] {X : Type*} [MulAction G X] {B : Set X} From a2bdc6e26b8b5e3ba7ca601ef2d952b17bd096ae Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Mon, 2 Dec 2024 05:47:14 +0000 Subject: [PATCH 458/829] =?UTF-8?q?feat(NumberTheory/NumberField/Completio?= =?UTF-8?q?n):=20API=20for=20`InfinitePlace=20=E2=84=9A`=20(#19644)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From the FLT project. When trying to do some calculations on adeles for FLT I realised that the completion of `ℚ` at its infinite place was not equal to `ℝ` and moreover we basically knew nothing about it. After this PR we at least know which rationals are in its open unit ball (which was what I needed). --- .../NumberTheory/NumberField/Completion.lean | 16 ++++++++++ .../NumberTheory/NumberField/Embeddings.lean | 31 +++++++++++++++++++ 2 files changed, 47 insertions(+) diff --git a/Mathlib/NumberTheory/NumberField/Completion.lean b/Mathlib/NumberTheory/NumberField/Completion.lean index 5f2aede3c8b65..901f318194033 100644 --- a/Mathlib/NumberTheory/NumberField/Completion.lean +++ b/Mathlib/NumberTheory/NumberField/Completion.lean @@ -70,9 +70,25 @@ instance : NormedField v.completion := letI := (WithAbs.isUniformInducing_of_comp v.norm_embedding_eq).completableTopField UniformSpace.Completion.instNormedFieldOfCompletableTopField (WithAbs v.1) +lemma norm_coe (x : WithAbs v.1) : + ‖(x : v.completion)‖ = v (WithAbs.equiv v.1 x) := + UniformSpace.Completion.norm_coe x + instance : Algebra K v.completion := inferInstanceAs <| Algebra (WithAbs v.1) v.1.completion +/-- The coercion from the rationals to its completion along an infinite place is `Rat.cast`. -/ +lemma WithAbs.ratCast_equiv (v : InfinitePlace ℚ) (x : WithAbs v.1) : + Rat.cast (WithAbs.equiv _ x) = (x : v.completion) := + (eq_ratCast (UniformSpace.Completion.coeRingHom.comp + (WithAbs.ringEquiv v.1).symm.toRingHom) x).symm + +lemma Rat.norm_infinitePlace_completion (v : InfinitePlace ℚ) (x : ℚ) : + ‖(x : v.completion)‖ = |x| := by + rw [← (WithAbs.equiv v.1).apply_symm_apply x, WithAbs.ratCast_equiv, + norm_coe, (WithAbs.equiv v.1).apply_symm_apply, + Rat.infinitePlace_apply] + /-- The completion of a number field at an infinite place is locally compact. -/ instance locallyCompactSpace : LocallyCompactSpace (v.completion) := AbsoluteValue.Completion.locallyCompactSpace v.norm_embedding_eq diff --git a/Mathlib/NumberTheory/NumberField/Embeddings.lean b/Mathlib/NumberTheory/NumberField/Embeddings.lean index 3a49b49ab4073..4b72199ed61ab 100644 --- a/Mathlib/NumberTheory/NumberField/Embeddings.lean +++ b/Mathlib/NumberTheory/NumberField/Embeddings.lean @@ -259,6 +259,13 @@ instance {K : Type*} [Field K] : FunLike (InfinitePlace K) K ℝ where coe w x := w.1 x coe_injective' _ _ h := Subtype.eq (AbsoluteValue.ext fun x => congr_fun h x) +lemma coe_apply {K : Type*} [Field K] (v : InfinitePlace K) (x : K) : + v x = v.1 x := rfl + +@[ext] +lemma ext {K : Type*} [Field K] (v₁ v₂ : InfinitePlace K) (h : ∀ k, v₁ k = v₂ k) : v₁ = v₂ := + Subtype.ext <| AbsoluteValue.ext h + instance : MonoidWithZeroHomClass (InfinitePlace K) K ℝ where map_mul w _ _ := w.1.map_mul _ _ map_one w := w.1.map_one @@ -1066,3 +1073,27 @@ theorem nrRealPlaces_eq_zero_of_two_lt (hk : 2 < k) (hζ : IsPrimitiveRoot ζ k) linarith end IsPrimitiveRoot + +/-! + +## The infinite place of the rationals. + +-/ + +namespace Rat + +open NumberField + +/-- The infinite place of `ℚ`, coming from the canonical map `ℚ → ℂ`. -/ +noncomputable def infinitePlace : InfinitePlace ℚ := .mk (Rat.castHom _) + +@[simp] +lemma infinitePlace_apply (v : InfinitePlace ℚ) (x : ℚ) : v x = |x| := by + rw [NumberField.InfinitePlace.coe_apply] + obtain ⟨_, _, rfl⟩ := v + simp + +instance : Subsingleton (InfinitePlace ℚ) where + allEq a b := by ext; simp + +end Rat From b90b953574df2fec78e73676597ac7aa35791dd0 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Mon, 2 Dec 2024 05:57:25 +0000 Subject: [PATCH 459/829] chore: update Mathlib dependencies 2024-12-02 (#19676) 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 0f235638e282f..5edf45f8e3379 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -68,7 +68,7 @@ "rev": "7805acf1864ba1a2e359f86a8f092ccf1438ad83", "name": "batteries", "manifestFile": "lake-manifest.json", - "inputRev": "v4.14.0", + "inputRev": "main", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover/lean4-cli", From e3df8ff92c33abe105e2a42a56b3646e9ae3ef62 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Mon, 2 Dec 2024 07:11:46 +0000 Subject: [PATCH 460/829] feat(Algebra): add two lemmas (#19661) Developed in conjunction with #19660, but found not to be needed. --- Mathlib/RingTheory/Adjoin/Basic.lean | 9 +++++++++ .../RingTheory/IntegralClosure/IsIntegral/Basic.lean | 11 ++++++++--- 2 files changed, 17 insertions(+), 3 deletions(-) diff --git a/Mathlib/RingTheory/Adjoin/Basic.lean b/Mathlib/RingTheory/Adjoin/Basic.lean index aa670454a746e..17423ca350195 100644 --- a/Mathlib/RingTheory/Adjoin/Basic.lean +++ b/Mathlib/RingTheory/Adjoin/Basic.lean @@ -322,6 +322,15 @@ theorem adjoin_adjoin_of_tower (s : Set A) : adjoin S (adjoin R s : Set A) = adj rw [this] exact subset_adjoin +theorem Subalgebra.restrictScalars_adjoin {s : Set A} : + (adjoin S s).restrictScalars R = (IsScalarTower.toAlgHom R S A).range ⊔ adjoin R s := by + refine le_antisymm (fun _ hx ↦ adjoin_induction + (fun x hx ↦ le_sup_right (α := Subalgebra R A) (subset_adjoin hx)) + (fun x ↦ le_sup_left (α := Subalgebra R A) ⟨x, rfl⟩) + (fun _ _ _ _ ↦ add_mem) (fun _ _ _ _ ↦ mul_mem) <| + (Subalgebra.mem_restrictScalars _).mp hx) (sup_le ?_ <| adjoin_le subset_adjoin) + rintro _ ⟨x, rfl⟩; exact algebraMap_mem (adjoin S s) x + @[simp] theorem adjoin_top {A} [Semiring A] [Algebra S A] (t : Set A) : adjoin (⊤ : Subalgebra R S) t = (adjoin S t).restrictScalars (⊤ : Subalgebra R S) := diff --git a/Mathlib/RingTheory/IntegralClosure/IsIntegral/Basic.lean b/Mathlib/RingTheory/IntegralClosure/IsIntegral/Basic.lean index e95d736928d15..c90e0cd0abbb5 100644 --- a/Mathlib/RingTheory/IntegralClosure/IsIntegral/Basic.lean +++ b/Mathlib/RingTheory/IntegralClosure/IsIntegral/Basic.lean @@ -100,9 +100,14 @@ theorem isIntegral_one [Algebra R B] : IsIntegral R (1 : B) := variable (f : R →+* S) theorem IsIntegral.of_pow [Algebra R B] {x : B} {n : ℕ} (hn : 0 < n) (hx : IsIntegral R <| x ^ n) : - IsIntegral R x := by - rcases hx with ⟨p, hmonic, heval⟩ - exact ⟨expand R n p, hmonic.expand hn, by rwa [← aeval_def, expand_aeval]⟩ + IsIntegral R x := + have ⟨p, hmonic, heval⟩ := hx + ⟨expand R n p, hmonic.expand hn, by rwa [← aeval_def, expand_aeval]⟩ + +theorem IsIntegral.of_aeval_monic {x : A} {p : R[X]} (monic : p.Monic) + (deg : p.natDegree ≠ 0) (hx : IsIntegral R (aeval x p)) : IsIntegral R x := + have ⟨p, hmonic, heval⟩ := hx + ⟨_, hmonic.comp monic deg, by rwa [eval₂_comp, ← aeval_def x]⟩ end From 720f1785d6c9ff4f2cfbcd007389d11ce9f1bac2 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 2 Dec 2024 07:36:55 +0000 Subject: [PATCH 461/829] chore: split Algebra.Submonoid.Membership (#19628) This file had two unrelated parts, one about big operators and one about other facts about membership. --- Mathlib.lean | 1 + Mathlib/Algebra/BigOperators/Associated.lean | 1 + Mathlib/Algebra/BigOperators/Finsupp.lean | 2 +- Mathlib/Algebra/Group/Subgroup/Finite.lean | 2 +- .../Algebra/Group/Submonoid/BigOperators.lean | 141 ++++++++++++++++++ .../Algebra/Group/Submonoid/Membership.lean | 119 +-------------- Mathlib/Algebra/Module/Submodule/Basic.lean | 2 +- Mathlib/Algebra/Module/Submodule/Lattice.lean | 1 + Mathlib/Algebra/Ring/Subsemiring/Basic.lean | 2 +- Mathlib/Data/DFinsupp/Submonoid.lean | 1 + Mathlib/GroupTheory/Perm/Sign.lean | 2 +- Mathlib/GroupTheory/QuotientGroup/Basic.lean | 1 + Mathlib/RingTheory/Localization/Defs.lean | 1 + .../RingTheory/NonUnitalSubring/Basic.lean | 1 + .../UniqueFactorizationDomain/Defs.lean | 1 + 15 files changed, 155 insertions(+), 123 deletions(-) create mode 100644 Mathlib/Algebra/Group/Submonoid/BigOperators.lean diff --git a/Mathlib.lean b/Mathlib.lean index 99e0381f97d9a..4e034e488a409 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -321,6 +321,7 @@ import Mathlib.Algebra.Group.Subgroup.Pointwise import Mathlib.Algebra.Group.Subgroup.ZPowers.Basic import Mathlib.Algebra.Group.Subgroup.ZPowers.Lemmas import Mathlib.Algebra.Group.Submonoid.Basic +import Mathlib.Algebra.Group.Submonoid.BigOperators import Mathlib.Algebra.Group.Submonoid.Defs import Mathlib.Algebra.Group.Submonoid.DistribMulAction import Mathlib.Algebra.Group.Submonoid.Membership diff --git a/Mathlib/Algebra/BigOperators/Associated.lean b/Mathlib/Algebra/BigOperators/Associated.lean index 6f770d8ad564c..e9978f3214640 100644 --- a/Mathlib/Algebra/BigOperators/Associated.lean +++ b/Mathlib/Algebra/BigOperators/Associated.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl, Jens Wagemaker, Anne Baanen -/ import Mathlib.Algebra.Associated.Basic import Mathlib.Algebra.BigOperators.Finsupp +import Mathlib.Algebra.Group.Submonoid.Membership /-! # Products of associated, prime, and irreducible elements. diff --git a/Mathlib/Algebra/BigOperators/Finsupp.lean b/Mathlib/Algebra/BigOperators/Finsupp.lean index dffd8f94b48be..dd91da8f00899 100644 --- a/Mathlib/Algebra/BigOperators/Finsupp.lean +++ b/Mathlib/Algebra/BigOperators/Finsupp.lean @@ -6,7 +6,7 @@ Authors: Kenny Lau import Mathlib.Algebra.BigOperators.Pi import Mathlib.Algebra.BigOperators.Ring import Mathlib.Algebra.BigOperators.Fin -import Mathlib.Algebra.Group.Submonoid.Membership +import Mathlib.Algebra.Group.Submonoid.BigOperators import Mathlib.Data.Finsupp.Fin import Mathlib.Data.Finsupp.Indicator diff --git a/Mathlib/Algebra/Group/Subgroup/Finite.lean b/Mathlib/Algebra/Group/Subgroup/Finite.lean index 6e5342a357f4e..db72bb5f2f470 100644 --- a/Mathlib/Algebra/Group/Subgroup/Finite.lean +++ b/Mathlib/Algebra/Group/Subgroup/Finite.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kexing Ying -/ import Mathlib.Algebra.Group.Subgroup.Basic -import Mathlib.Algebra.Group.Submonoid.Membership +import Mathlib.Algebra.Group.Submonoid.BigOperators import Mathlib.Data.Finite.Card /-! diff --git a/Mathlib/Algebra/Group/Submonoid/BigOperators.lean b/Mathlib/Algebra/Group/Submonoid/BigOperators.lean new file mode 100644 index 0000000000000..5d673c2469b14 --- /dev/null +++ b/Mathlib/Algebra/Group/Submonoid/BigOperators.lean @@ -0,0 +1,141 @@ +/- +Copyright (c) 2018 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Kenny Lau, Johan Commelin, Mario Carneiro, Kevin Buzzard, +Amelia Livingston, Yury Kudryashov +-/ +import Mathlib.Algebra.Group.Submonoid.Defs +import Mathlib.Data.Finset.NoncommProd + +/-! +# Submonoids: membership criteria for products and sums + +In this file we prove various facts about membership in a submonoid: + +* `list_prod_mem`, `multiset_prod_mem`, `prod_mem`: if each element of a collection belongs + to a multiplicative submonoid, then so does their product; +* `list_sum_mem`, `multiset_sum_mem`, `sum_mem`: if each element of a collection belongs + to an additive submonoid, then so does their sum; + +## Tags +submonoid, submonoids +-/ + +-- We don't need ordered structures to establish basic membership facts for submonoids +assert_not_exists OrderedSemiring + +variable {M A B : Type*} + +variable [Monoid M] [SetLike B M] [SubmonoidClass B M] {S : B} + +namespace SubmonoidClass + +@[to_additive (attr := norm_cast, simp)] +theorem coe_list_prod (l : List S) : (l.prod : M) = (l.map (↑)).prod := + map_list_prod (SubmonoidClass.subtype S : _ →* M) l + +@[to_additive (attr := norm_cast, simp)] +theorem coe_multiset_prod {M} [CommMonoid M] [SetLike B M] [SubmonoidClass B M] (m : Multiset S) : + (m.prod : M) = (m.map (↑)).prod := + (SubmonoidClass.subtype S : _ →* M).map_multiset_prod m + +@[to_additive (attr := norm_cast, simp)] +theorem coe_finset_prod {ι M} [CommMonoid M] [SetLike B M] [SubmonoidClass B M] (f : ι → S) + (s : Finset ι) : ↑(∏ i ∈ s, f i) = (∏ i ∈ s, f i : M) := + map_prod (SubmonoidClass.subtype S) f s + +end SubmonoidClass + +open SubmonoidClass + +/-- Product of a list of elements in a submonoid is in the submonoid. -/ +@[to_additive "Sum of a list of elements in an `AddSubmonoid` is in the `AddSubmonoid`."] +theorem list_prod_mem {l : List M} (hl : ∀ x ∈ l, x ∈ S) : l.prod ∈ S := by + lift l to List S using hl + rw [← coe_list_prod] + exact l.prod.coe_prop + +/-- Product of a multiset of elements in a submonoid of a `CommMonoid` is in the submonoid. -/ +@[to_additive + "Sum of a multiset of elements in an `AddSubmonoid` of an `AddCommMonoid` is + in the `AddSubmonoid`."] +theorem multiset_prod_mem {M} [CommMonoid M] [SetLike B M] [SubmonoidClass B M] (m : Multiset M) + (hm : ∀ a ∈ m, a ∈ S) : m.prod ∈ S := by + lift m to Multiset S using hm + rw [← coe_multiset_prod] + exact m.prod.coe_prop + +/-- Product of elements of a submonoid of a `CommMonoid` indexed by a `Finset` is in the + submonoid. -/ +@[to_additive + "Sum of elements in an `AddSubmonoid` of an `AddCommMonoid` indexed by a `Finset` + is in the `AddSubmonoid`."] +theorem prod_mem {M : Type*} [CommMonoid M] [SetLike B M] [SubmonoidClass B M] {ι : Type*} + {t : Finset ι} {f : ι → M} (h : ∀ c ∈ t, f c ∈ S) : (∏ c ∈ t, f c) ∈ S := + multiset_prod_mem (t.1.map f) fun _x hx => + let ⟨i, hi, hix⟩ := Multiset.mem_map.1 hx + hix ▸ h i hi + +namespace Submonoid + +variable (s : Submonoid M) + +@[to_additive (attr := norm_cast)] +theorem coe_list_prod (l : List s) : (l.prod : M) = (l.map (↑)).prod := + map_list_prod s.subtype l + +@[to_additive (attr := norm_cast)] +theorem coe_multiset_prod {M} [CommMonoid M] (S : Submonoid M) (m : Multiset S) : + (m.prod : M) = (m.map (↑)).prod := + S.subtype.map_multiset_prod m + +@[to_additive (attr := norm_cast)] +theorem coe_finset_prod {ι M} [CommMonoid M] (S : Submonoid M) (f : ι → S) (s : Finset ι) : + ↑(∏ i ∈ s, f i) = (∏ i ∈ s, f i : M) := + map_prod S.subtype f s + +/-- Product of a list of elements in a submonoid is in the submonoid. -/ +@[to_additive "Sum of a list of elements in an `AddSubmonoid` is in the `AddSubmonoid`."] +theorem list_prod_mem {l : List M} (hl : ∀ x ∈ l, x ∈ s) : l.prod ∈ s := by + lift l to List s using hl + rw [← coe_list_prod] + exact l.prod.coe_prop + +/-- Product of a multiset of elements in a submonoid of a `CommMonoid` is in the submonoid. -/ +@[to_additive + "Sum of a multiset of elements in an `AddSubmonoid` of an `AddCommMonoid` is + in the `AddSubmonoid`."] +theorem multiset_prod_mem {M} [CommMonoid M] (S : Submonoid M) (m : Multiset M) + (hm : ∀ a ∈ m, a ∈ S) : m.prod ∈ S := by + lift m to Multiset S using hm + rw [← coe_multiset_prod] + exact m.prod.coe_prop + +@[to_additive] +theorem multiset_noncommProd_mem (S : Submonoid M) (m : Multiset M) (comm) (h : ∀ x ∈ m, x ∈ S) : + m.noncommProd comm ∈ S := by + induction m using Quotient.inductionOn with | h l => ?_ + simp only [Multiset.quot_mk_to_coe, Multiset.noncommProd_coe] + exact Submonoid.list_prod_mem _ h + +/-- Product of elements of a submonoid of a `CommMonoid` indexed by a `Finset` is in the + submonoid. -/ +@[to_additive + "Sum of elements in an `AddSubmonoid` of an `AddCommMonoid` indexed by a `Finset` + is in the `AddSubmonoid`."] +theorem prod_mem {M : Type*} [CommMonoid M] (S : Submonoid M) {ι : Type*} {t : Finset ι} + {f : ι → M} (h : ∀ c ∈ t, f c ∈ S) : (∏ c ∈ t, f c) ∈ S := + S.multiset_prod_mem (t.1.map f) fun _ hx => + let ⟨i, hi, hix⟩ := Multiset.mem_map.1 hx + hix ▸ h i hi + +@[to_additive] +theorem noncommProd_mem (S : Submonoid M) {ι : Type*} (t : Finset ι) (f : ι → M) (comm) + (h : ∀ c ∈ t, f c ∈ S) : t.noncommProd f comm ∈ S := by + apply multiset_noncommProd_mem + intro y + rw [Multiset.mem_map] + rintro ⟨x, ⟨hx, rfl⟩⟩ + exact h x hx + +end Submonoid diff --git a/Mathlib/Algebra/Group/Submonoid/Membership.lean b/Mathlib/Algebra/Group/Submonoid/Membership.lean index 2eeb9e1f37c14..82f50808f2140 100644 --- a/Mathlib/Algebra/Group/Submonoid/Membership.lean +++ b/Mathlib/Algebra/Group/Submonoid/Membership.lean @@ -10,19 +10,14 @@ import Mathlib.Algebra.Group.Submonoid.Operations import Mathlib.Algebra.GroupWithZero.Divisibility import Mathlib.Algebra.Ring.Idempotents import Mathlib.Algebra.Ring.Int.Defs -import Mathlib.Data.Finset.NoncommProd +import Mathlib.Data.Fintype.Card import Mathlib.Data.Nat.Cast.Basic -import Mathlib.Util.AssertExists /-! # Submonoids: membership criteria In this file we prove various facts about membership in a submonoid: -* `list_prod_mem`, `multiset_prod_mem`, `prod_mem`: if each element of a collection belongs - to a multiplicative submonoid, then so does their product; -* `list_sum_mem`, `multiset_sum_mem`, `sum_mem`: if each element of a collection belongs - to an additive submonoid, then so does their sum; * `pow_mem`, `nsmul_mem`: if `x ∈ S` where `S` is a multiplicative (resp., additive) submonoid and `n` is a natural number, then `x^n` (resp., `n • x`) belongs to `S`; * `mem_iSup_of_directed`, `coe_iSup_of_directed`, `mem_sSup_of_directedOn`, @@ -46,118 +41,6 @@ section Assoc variable [Monoid M] [SetLike B M] [SubmonoidClass B M] {S : B} -namespace SubmonoidClass - -@[to_additive (attr := norm_cast, simp)] -theorem coe_list_prod (l : List S) : (l.prod : M) = (l.map (↑)).prod := - map_list_prod (SubmonoidClass.subtype S : _ →* M) l - -@[to_additive (attr := norm_cast, simp)] -theorem coe_multiset_prod {M} [CommMonoid M] [SetLike B M] [SubmonoidClass B M] (m : Multiset S) : - (m.prod : M) = (m.map (↑)).prod := - (SubmonoidClass.subtype S : _ →* M).map_multiset_prod m - -@[to_additive (attr := norm_cast, simp)] -theorem coe_finset_prod {ι M} [CommMonoid M] [SetLike B M] [SubmonoidClass B M] (f : ι → S) - (s : Finset ι) : ↑(∏ i ∈ s, f i) = (∏ i ∈ s, f i : M) := - map_prod (SubmonoidClass.subtype S) f s - -end SubmonoidClass - -open SubmonoidClass - -/-- Product of a list of elements in a submonoid is in the submonoid. -/ -@[to_additive "Sum of a list of elements in an `AddSubmonoid` is in the `AddSubmonoid`."] -theorem list_prod_mem {l : List M} (hl : ∀ x ∈ l, x ∈ S) : l.prod ∈ S := by - lift l to List S using hl - rw [← coe_list_prod] - exact l.prod.coe_prop - -/-- Product of a multiset of elements in a submonoid of a `CommMonoid` is in the submonoid. -/ -@[to_additive - "Sum of a multiset of elements in an `AddSubmonoid` of an `AddCommMonoid` is - in the `AddSubmonoid`."] -theorem multiset_prod_mem {M} [CommMonoid M] [SetLike B M] [SubmonoidClass B M] (m : Multiset M) - (hm : ∀ a ∈ m, a ∈ S) : m.prod ∈ S := by - lift m to Multiset S using hm - rw [← coe_multiset_prod] - exact m.prod.coe_prop - -/-- Product of elements of a submonoid of a `CommMonoid` indexed by a `Finset` is in the - submonoid. -/ -@[to_additive - "Sum of elements in an `AddSubmonoid` of an `AddCommMonoid` indexed by a `Finset` - is in the `AddSubmonoid`."] -theorem prod_mem {M : Type*} [CommMonoid M] [SetLike B M] [SubmonoidClass B M] {ι : Type*} - {t : Finset ι} {f : ι → M} (h : ∀ c ∈ t, f c ∈ S) : (∏ c ∈ t, f c) ∈ S := - multiset_prod_mem (t.1.map f) fun _x hx => - let ⟨i, hi, hix⟩ := Multiset.mem_map.1 hx - hix ▸ h i hi - -namespace Submonoid - -variable (s : Submonoid M) - -@[to_additive (attr := norm_cast)] -theorem coe_list_prod (l : List s) : (l.prod : M) = (l.map (↑)).prod := - map_list_prod s.subtype l - -@[to_additive (attr := norm_cast)] -theorem coe_multiset_prod {M} [CommMonoid M] (S : Submonoid M) (m : Multiset S) : - (m.prod : M) = (m.map (↑)).prod := - S.subtype.map_multiset_prod m - -@[to_additive (attr := norm_cast)] -theorem coe_finset_prod {ι M} [CommMonoid M] (S : Submonoid M) (f : ι → S) (s : Finset ι) : - ↑(∏ i ∈ s, f i) = (∏ i ∈ s, f i : M) := - map_prod S.subtype f s - -/-- Product of a list of elements in a submonoid is in the submonoid. -/ -@[to_additive "Sum of a list of elements in an `AddSubmonoid` is in the `AddSubmonoid`."] -theorem list_prod_mem {l : List M} (hl : ∀ x ∈ l, x ∈ s) : l.prod ∈ s := by - lift l to List s using hl - rw [← coe_list_prod] - exact l.prod.coe_prop - -/-- Product of a multiset of elements in a submonoid of a `CommMonoid` is in the submonoid. -/ -@[to_additive - "Sum of a multiset of elements in an `AddSubmonoid` of an `AddCommMonoid` is - in the `AddSubmonoid`."] -theorem multiset_prod_mem {M} [CommMonoid M] (S : Submonoid M) (m : Multiset M) - (hm : ∀ a ∈ m, a ∈ S) : m.prod ∈ S := by - lift m to Multiset S using hm - rw [← coe_multiset_prod] - exact m.prod.coe_prop - -@[to_additive] -theorem multiset_noncommProd_mem (S : Submonoid M) (m : Multiset M) (comm) (h : ∀ x ∈ m, x ∈ S) : - m.noncommProd comm ∈ S := by - induction m using Quotient.inductionOn with | h l => ?_ - simp only [Multiset.quot_mk_to_coe, Multiset.noncommProd_coe] - exact Submonoid.list_prod_mem _ h - -/-- Product of elements of a submonoid of a `CommMonoid` indexed by a `Finset` is in the - submonoid. -/ -@[to_additive - "Sum of elements in an `AddSubmonoid` of an `AddCommMonoid` indexed by a `Finset` - is in the `AddSubmonoid`."] -theorem prod_mem {M : Type*} [CommMonoid M] (S : Submonoid M) {ι : Type*} {t : Finset ι} - {f : ι → M} (h : ∀ c ∈ t, f c ∈ S) : (∏ c ∈ t, f c) ∈ S := - S.multiset_prod_mem (t.1.map f) fun _ hx => - let ⟨i, hi, hix⟩ := Multiset.mem_map.1 hx - hix ▸ h i hi - -@[to_additive] -theorem noncommProd_mem (S : Submonoid M) {ι : Type*} (t : Finset ι) (f : ι → M) (comm) - (h : ∀ c ∈ t, f c ∈ S) : t.noncommProd f comm ∈ S := by - apply multiset_noncommProd_mem - intro y - rw [Multiset.mem_map] - rintro ⟨x, ⟨hx, rfl⟩⟩ - exact h x hx - -end Submonoid - end Assoc section NonAssoc diff --git a/Mathlib/Algebra/Module/Submodule/Basic.lean b/Mathlib/Algebra/Module/Submodule/Basic.lean index 05f1f2a66fbab..e6ab0afbfd51c 100644 --- a/Mathlib/Algebra/Module/Submodule/Basic.lean +++ b/Mathlib/Algebra/Module/Submodule/Basic.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 -/ import Mathlib.Algebra.Field.Defs -import Mathlib.Algebra.Group.Submonoid.Membership +import Mathlib.Algebra.Group.Submonoid.BigOperators import Mathlib.Algebra.Module.Submodule.Defs import Mathlib.Algebra.NoZeroSMulDivisors.Defs import Mathlib.GroupTheory.GroupAction.SubMulAction diff --git a/Mathlib/Algebra/Module/Submodule/Lattice.lean b/Mathlib/Algebra/Module/Submodule/Lattice.lean index 76744e1fa2be2..364a656683578 100644 --- a/Mathlib/Algebra/Module/Submodule/Lattice.lean +++ b/Mathlib/Algebra/Module/Submodule/Lattice.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl, Mario Carneiro, Kevin Buzzard, Yury Kudryashov -/ import Mathlib.Algebra.Group.Subgroup.Lattice import Mathlib.Algebra.Group.Submonoid.Membership +import Mathlib.Algebra.Group.Submonoid.BigOperators import Mathlib.Algebra.Module.Submodule.Defs import Mathlib.Algebra.Module.Equiv.Defs import Mathlib.Algebra.PUnitInstances.Module diff --git a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean index 5a0b8a674f389..65c4d305a67a1 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import Mathlib.Algebra.Group.Submonoid.Membership +import Mathlib.Algebra.Group.Submonoid.BigOperators import Mathlib.Algebra.Module.RingHom import Mathlib.Algebra.Ring.Action.Subobjects import Mathlib.Algebra.Ring.Equiv diff --git a/Mathlib/Data/DFinsupp/Submonoid.lean b/Mathlib/Data/DFinsupp/Submonoid.lean index 1bcd75a448f19..4af6b95649e4d 100644 --- a/Mathlib/Data/DFinsupp/Submonoid.lean +++ b/Mathlib/Data/DFinsupp/Submonoid.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Kenny Lau -/ +import Mathlib.Algebra.Group.Submonoid.BigOperators import Mathlib.Algebra.Group.Submonoid.Membership import Mathlib.Data.DFinsupp.BigOperators import Mathlib.Order.ConditionallyCompleteLattice.Basic diff --git a/Mathlib/GroupTheory/Perm/Sign.lean b/Mathlib/GroupTheory/Perm/Sign.lean index bfb55453b7c78..d8fe100a71871 100644 --- a/Mathlib/GroupTheory/Perm/Sign.lean +++ b/Mathlib/GroupTheory/Perm/Sign.lean @@ -5,7 +5,7 @@ Authors: Chris Hughes -/ import Mathlib.Algebra.Group.Conj import Mathlib.Algebra.Group.Subgroup.Lattice -import Mathlib.Algebra.Group.Submonoid.Membership +import Mathlib.Algebra.Group.Submonoid.BigOperators import Mathlib.Algebra.Ring.Int.Units import Mathlib.Data.Finset.Fin import Mathlib.Data.Finset.Sort diff --git a/Mathlib/GroupTheory/QuotientGroup/Basic.lean b/Mathlib/GroupTheory/QuotientGroup/Basic.lean index 3c1822db6828d..58622e2d53e83 100644 --- a/Mathlib/GroupTheory/QuotientGroup/Basic.lean +++ b/Mathlib/GroupTheory/QuotientGroup/Basic.lean @@ -10,6 +10,7 @@ import Mathlib.Data.Int.Cast.Lemmas import Mathlib.GroupTheory.Congruence.Hom import Mathlib.GroupTheory.Coset.Basic import Mathlib.GroupTheory.QuotientGroup.Defs +import Mathlib.Algebra.BigOperators.Group.Finset /-! # Quotients of groups by normal subgroups diff --git a/Mathlib/RingTheory/Localization/Defs.lean b/Mathlib/RingTheory/Localization/Defs.lean index 60509638366d3..afe729f268f40 100644 --- a/Mathlib/RingTheory/Localization/Defs.lean +++ b/Mathlib/RingTheory/Localization/Defs.lean @@ -8,6 +8,7 @@ import Mathlib.GroupTheory.MonoidLocalization.MonoidWithZero import Mathlib.RingTheory.OreLocalization.Ring import Mathlib.Tactic.ApplyFun import Mathlib.Tactic.Ring +import Mathlib.Algebra.BigOperators.Group.Finset /-! # Localizations of commutative rings diff --git a/Mathlib/RingTheory/NonUnitalSubring/Basic.lean b/Mathlib/RingTheory/NonUnitalSubring/Basic.lean index e376982889514..e09499939d3a0 100644 --- a/Mathlib/RingTheory/NonUnitalSubring/Basic.lean +++ b/Mathlib/RingTheory/NonUnitalSubring/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ import Mathlib.Algebra.Group.Subgroup.Basic +import Mathlib.Algebra.Group.Submonoid.BigOperators import Mathlib.GroupTheory.Subsemigroup.Center import Mathlib.RingTheory.NonUnitalSubring.Defs import Mathlib.RingTheory.NonUnitalSubsemiring.Basic diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Defs.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Defs.lean index 121cf52e3a622..28fef6a3998c7 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain/Defs.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Defs.lean @@ -6,6 +6,7 @@ Authors: Johannes Hölzl, Jens Wagemaker, Aaron Anderson import Mathlib.Algebra.Associated.Basic import Mathlib.Algebra.BigOperators.Group.Multiset import Mathlib.Algebra.Group.Submonoid.Membership +import Mathlib.Algebra.Group.Submonoid.BigOperators import Mathlib.Order.WellFounded /-! From 675c28ea285896bbe83cb32f8a0bca7e481be614 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Mon, 2 Dec 2024 07:36:56 +0000 Subject: [PATCH 462/829] feat(Polynomial): new lemmas and noncommutative generalizations (#19665) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In `Algebra/Polynomial/Eval/Degree.lean`: two new lemmas about the coefficients and zero-ness of the composition of a polynomial with `C r * X`. Noncommutative generalizations: `Algebra/Polynomial/Degree/Lemmas.lean`: two lemmas `RingTheory/Polynomial/Tower.lean`: one lemma `RingTheory/Polynomial/IntegralNormalization.lean`: add `of_commute` versions of two lemmas similar to [Polynomial.scaleRoots_eval₂_mul_of_commute](https://leanprover-community.github.io/mathlib4_docs/Mathlib/RingTheory/Polynomial/ScaleRoots.html#Polynomial.scaleRoots_eval%E2%82%82_mul_of_commute) and generalize one lemma. --- Mathlib/Algebra/Polynomial/Degree/Lemmas.lean | 14 ++----- Mathlib/Algebra/Polynomial/Eval/Degree.lean | 16 ++++++++ .../Polynomial/IntegralNormalization.lean | 40 +++++++++++-------- Mathlib/RingTheory/Polynomial/Tower.lean | 6 +-- 4 files changed, 46 insertions(+), 30 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean b/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean index 5e474102b75cc..6f001cb5e20e7 100644 --- a/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean +++ b/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean @@ -369,20 +369,16 @@ theorem leadingCoeff_comp (hq : natDegree q ≠ 0) : end NoZeroDivisors -section CommRing -variable [CommRing R] {p q : R[X]} - -@[simp] lemma comp_neg_X_leadingCoeff_eq (p : R[X]) : +@[simp] lemma comp_neg_X_leadingCoeff_eq [Ring R] (p : R[X]) : (p.comp (-X)).leadingCoeff = (-1) ^ p.natDegree * p.leadingCoeff := by nontriviality R by_cases h : p = 0 · simp [h] rw [Polynomial.leadingCoeff, natDegree_comp_eq_of_mul_ne_zero, coeff_comp_degree_mul_degree] <;> - simp [mul_comm, h] - -variable [IsDomain R] + simp [((Commute.neg_one_left _).pow_left _).eq, h] -lemma comp_eq_zero_iff : p.comp q = 0 ↔ p = 0 ∨ p.eval (q.coeff 0) = 0 ∧ q = C (q.coeff 0) := by +lemma comp_eq_zero_iff [Semiring R] [NoZeroDivisors R] {p q : R[X]} : + p.comp q = 0 ↔ p = 0 ∨ p.eval (q.coeff 0) = 0 ∧ q = C (q.coeff 0) := by refine ⟨fun h ↦ ?_, Or.rec (fun h ↦ by simp [h]) fun h ↦ by rw [h.2, comp_C, h.1, C_0]⟩ have key : p.natDegree = 0 ∨ q.natDegree = 0 := by rw [← mul_eq_zero, ← natDegree_comp, h, natDegree_zero] @@ -392,8 +388,6 @@ lemma comp_eq_zero_iff : p.comp q = 0 ↔ p = 0 ∨ p.eval (q.coeff 0) = 0 ∧ q · rw [key, comp_C, C_eq_zero] at h exact Or.inr ⟨h, key⟩ -end CommRing - section DivisionRing variable {K : Type*} [DivisionRing K] diff --git a/Mathlib/Algebra/Polynomial/Eval/Degree.lean b/Mathlib/Algebra/Polynomial/Eval/Degree.lean index 41edbc137ec86..4105940c90547 100644 --- a/Mathlib/Algebra/Polynomial/Eval/Degree.lean +++ b/Mathlib/Algebra/Polynomial/Eval/Degree.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Johannes Hölzl, Kim Morrison, Jens Wagemaker -/ +import Mathlib.Algebra.GroupWithZero.NonZeroDivisors import Mathlib.Algebra.Polynomial.Degree.Support import Mathlib.Algebra.Polynomial.Degree.Units import Mathlib.Algebra.Polynomial.Eval.Coeff @@ -107,6 +108,21 @@ theorem coeff_comp_degree_mul_degree (hqd0 : natDegree q ≠ 0) : case h₁ => simp +contextual +@[simp] lemma comp_C_mul_X_coeff {r : R} {n : ℕ} : + (p.comp <| C r * X).coeff n = p.coeff n * r ^ n := by + simp_rw [comp, eval₂_eq_sum_range, (commute_X _).symm.mul_pow, + ← C_pow, finset_sum_coeff, coeff_C_mul, coeff_X_pow] + rw [Finset.sum_eq_single n _ fun h ↦ ?_, if_pos rfl, mul_one] + · intro b _ h; simp_rw [if_neg h.symm, mul_zero] + · rw [coeff_eq_zero_of_natDegree_lt, zero_mul] + rwa [Finset.mem_range_succ_iff, not_le] at h + +lemma comp_C_mul_X_eq_zero_iff {r : R} (hr : r ∈ nonZeroDivisors R) : + p.comp (C r * X) = 0 ↔ p = 0 := by + simp_rw [ext_iff] + refine forall_congr' fun n ↦ ?_ + rw [comp_C_mul_X_coeff, coeff_zero, mul_right_mem_nonZeroDivisors_eq_zero_iff (pow_mem hr _)] + end Comp section Map diff --git a/Mathlib/RingTheory/Polynomial/IntegralNormalization.lean b/Mathlib/RingTheory/Polynomial/IntegralNormalization.lean index 046c30220049a..ee4f04ec120b8 100644 --- a/Mathlib/RingTheory/Polynomial/IntegralNormalization.lean +++ b/Mathlib/RingTheory/Polynomial/IntegralNormalization.lean @@ -115,24 +115,30 @@ theorem integralNormalization_degree : (integralNormalization p).degree = p.degr · rw [← degree_scaleRoots, ← integralNormalization_mul_C_leadingCoeff] exact (degree_mul_le _ _).trans (add_le_of_nonpos_right degree_C_le) -variable [CommSemiring S] +variable {A : Type*} [CommSemiring S] [Semiring A] theorem leadingCoeff_smul_integralNormalization (p : S[X]) : p.leadingCoeff • integralNormalization p = scaleRoots p p.leadingCoeff := by rw [Algebra.smul_def, algebraMap_eq, mul_comm, integralNormalization_mul_C_leadingCoeff] -theorem integralNormalization_eval₂_leadingCoeff_mul (h : 1 ≤ p.natDegree) (f : R →+* S) (x : S) : +theorem integralNormalization_eval₂_leadingCoeff_mul_of_commute (h : 1 ≤ p.natDegree) (f : R →+* A) + (x : A) (h₁ : Commute (f p.leadingCoeff) x) (h₂ : ∀ {r r'}, Commute (f r) (f r')) : (integralNormalization p).eval₂ f (f p.leadingCoeff * x) = f p.leadingCoeff ^ (p.natDegree - 1) * p.eval₂ f x := by rw [eval₂_eq_sum_range, eval₂_eq_sum_range, Finset.mul_sum] apply Finset.sum_congr · rw [natDegree_eq_of_degree_eq p.integralNormalization_degree] intro n _hn - rw [mul_pow, ← mul_assoc, ← f.map_pow, ← f.map_mul, - integralNormalization_coeff_mul_leadingCoeff_pow _ h, f.map_mul, f.map_pow] - ring + rw [h₁.mul_pow, ← mul_assoc, ← f.map_pow, ← f.map_mul, + integralNormalization_coeff_mul_leadingCoeff_pow _ h, f.map_mul, h₂.eq, f.map_pow, mul_assoc] -theorem integralNormalization_eval₂_eq_zero {p : R[X]} (f : R →+* S) {z : S} (hz : eval₂ f z p = 0) +theorem integralNormalization_eval₂_leadingCoeff_mul (h : 1 ≤ p.natDegree) (f : R →+* S) (x : S) : + (integralNormalization p).eval₂ f (f p.leadingCoeff * x) = + f p.leadingCoeff ^ (p.natDegree - 1) * p.eval₂ f x := + integralNormalization_eval₂_leadingCoeff_mul_of_commute h _ _ (.all _ _) (.all _ _) + +theorem integralNormalization_eval₂_eq_zero_of_commute {p : R[X]} (f : R →+* A) {z : A} + (hz : eval₂ f z p = 0) (h₁ : Commute (f p.leadingCoeff) z) (h₂ : ∀ {r r'}, Commute (f r) (f r')) (inj : ∀ x : R, f x = 0 → x = 0) : eval₂ f (f p.leadingCoeff * z) (integralNormalization p) = 0 := by obtain (h | h) := p.natDegree.eq_zero_or_pos @@ -141,20 +147,20 @@ theorem integralNormalization_eval₂_eq_zero {p : R[X]} (f : R →+* S) {z : S} simp [h0] · rw [eq_C_of_natDegree_eq_zero h, eval₂_C] at hz exact absurd (inj _ hz) h0 - · rw [integralNormalization_eval₂_leadingCoeff_mul h, hz, mul_zero] + · rw [integralNormalization_eval₂_leadingCoeff_mul_of_commute h _ _ h₁ h₂, hz, mul_zero] -end Semiring - -section CommSemiring - -variable [CommSemiring R] [CommSemiring S] +theorem integralNormalization_eval₂_eq_zero {p : R[X]} (f : R →+* S) {z : S} (hz : eval₂ f z p = 0) + (inj : ∀ x : R, f x = 0 → x = 0) : + eval₂ f (f p.leadingCoeff * z) (integralNormalization p) = 0 := + integralNormalization_eval₂_eq_zero_of_commute _ hz (.all _ _) (.all _ _) inj -theorem integralNormalization_aeval_eq_zero [Algebra R S] {f : R[X]} {z : S} (hz : aeval z f = 0) - (inj : ∀ x : R, algebraMap R S x = 0 → x = 0) : - aeval (algebraMap R S f.leadingCoeff * z) (integralNormalization f) = 0 := - integralNormalization_eval₂_eq_zero (algebraMap R S) hz inj +theorem integralNormalization_aeval_eq_zero [Algebra S A] {f : S[X]} {z : A} (hz : aeval z f = 0) + (inj : ∀ x : S, algebraMap S A x = 0 → x = 0) : + aeval (algebraMap S A f.leadingCoeff * z) (integralNormalization f) = 0 := + integralNormalization_eval₂_eq_zero_of_commute (algebraMap S A) hz + (Algebra.commute_algebraMap_left _ _) (.map (.all _ _) _) inj -end CommSemiring +end Semiring section IsCancelMulZero diff --git a/Mathlib/RingTheory/Polynomial/Tower.lean b/Mathlib/RingTheory/Polynomial/Tower.lean index 5e49599ddbda9..3177903b37969 100644 --- a/Mathlib/RingTheory/Polynomial/Tower.lean +++ b/Mathlib/RingTheory/Polynomial/Tower.lean @@ -36,9 +36,9 @@ theorem aeval_map_algebraMap (x : B) (p : R[X]) : aeval x (map (algebraMap R A) rw [aeval_def, aeval_def, eval₂_map, IsScalarTower.algebraMap_eq R A B] @[simp] -lemma eval_map_algebraMap (P : R[X]) (a : A) : - (map (algebraMap R A) P).eval a = aeval a P := by - rw [← aeval_map_algebraMap (A := A), coe_aeval_eq_eval] +lemma eval_map_algebraMap (P : R[X]) (b : B) : + (map (algebraMap R B) P).eval b = aeval b P := by + rw [aeval_def, eval_map] end Semiring From 17ec70ce0d0fcc06b537498327827b4d5a20ee78 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Mon, 2 Dec 2024 11:22:54 +0000 Subject: [PATCH 463/829] chore(Algebra/Polynomial): don't import `Ideal` when defining `Polynomial.roots` (#19376) This PR moves around a few definitions in order to free the definition of `Polynomial.roots` of the knowledge of `Ideal`. This was mostly motivated by writing `assert_not_exists Ideal` in a few files and being amazed that it did exist. In particular, the following changes are made: * results involving ideals in `Algebra/Module/Submodule/Pointwise.lean` go to `RingTheory/Ideal/Operations.lean` (the first place that they are used) * results on ideal membership in `RingTheory/Localization/Defs.lean` go to `RingTheory/Localization/Ideal.lean` (idem) * results on adjoining polynomials to a ring in `Algebra/Polynomial/AlgebraMap.lean` go to the new file `RingTheory/Adjoin/Polynomial.lean` (I couldn't find an obvious place to put this: `RingTheory/Adjoin/Basic.lean` does not know polynomials) * results on ideal membership in `Algebra/Polynomial/Div.lean`, `Algebra/Polynomial/RingDivision.lean` and a few from `RingTheory/Polynomial/Basic.lean` go to the new file `RingTheory/Polynomial/Ideal.lean` (merging this with `RingTheory/Polynomial/Basic.lean` is an okay option but would increase the downstream imports a lot) --- Mathlib.lean | 2 + .../Algebra/Module/Submodule/Pointwise.lean | 33 +------- Mathlib/Algebra/Polynomial/AlgebraMap.lean | 36 +-------- Mathlib/Algebra/Polynomial/Derivation.lean | 1 + Mathlib/Algebra/Polynomial/Div.lean | 18 +---- .../Algebra/Polynomial/PartialFractions.lean | 1 + Mathlib/Algebra/Polynomial/RingDivision.lean | 9 +-- Mathlib/Algebra/Polynomial/Roots.lean | 4 +- Mathlib/Algebra/Polynomial/Splits.lean | 1 + Mathlib/Analysis/Normed/Field/Basic.lean | 1 + .../FieldTheory/IntermediateField/Basic.lean | 1 + Mathlib/LinearAlgebra/Matrix/ToLin.lean | 1 + Mathlib/RingTheory/Adjoin/Basic.lean | 1 + Mathlib/RingTheory/Adjoin/Polynomial.lean | 67 ++++++++++++++++ .../RingTheory/Algebraic/MvPolynomial.lean | 1 + Mathlib/RingTheory/FiniteType.lean | 1 + .../RingTheory/HahnSeries/Multiplication.lean | 1 + Mathlib/RingTheory/Ideal/Operations.lean | 28 +++++++ .../IntegralClosure/IsIntegral/Basic.lean | 1 + Mathlib/RingTheory/IsTensorProduct.lean | 3 +- Mathlib/RingTheory/Localization/Basic.lean | 11 +-- .../RingTheory/Localization/FractionRing.lean | 1 + Mathlib/RingTheory/Localization/Ideal.lean | 12 ++- Mathlib/RingTheory/Polynomial/Basic.lean | 31 -------- .../Polynomial/Eisenstein/Basic.lean | 1 + Mathlib/RingTheory/Polynomial/Ideal.lean | 76 +++++++++++++++++++ .../Polynomial/IntegralNormalization.lean | 1 + Mathlib/RingTheory/Polynomial/Quotient.lean | 1 + 28 files changed, 216 insertions(+), 129 deletions(-) create mode 100644 Mathlib/RingTheory/Adjoin/Polynomial.lean create mode 100644 Mathlib/RingTheory/Polynomial/Ideal.lean diff --git a/Mathlib.lean b/Mathlib.lean index 4e034e488a409..fccdcac503167 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4215,6 +4215,7 @@ import Mathlib.RingTheory.Adjoin.Basic import Mathlib.RingTheory.Adjoin.Dimension import Mathlib.RingTheory.Adjoin.FG import Mathlib.RingTheory.Adjoin.Field +import Mathlib.RingTheory.Adjoin.Polynomial import Mathlib.RingTheory.Adjoin.PowerBasis import Mathlib.RingTheory.Adjoin.Tower import Mathlib.RingTheory.AdjoinRoot @@ -4474,6 +4475,7 @@ import Mathlib.RingTheory.Polynomial.Eisenstein.IsIntegral import Mathlib.RingTheory.Polynomial.GaussLemma import Mathlib.RingTheory.Polynomial.Hermite.Basic import Mathlib.RingTheory.Polynomial.Hermite.Gaussian +import Mathlib.RingTheory.Polynomial.Ideal import Mathlib.RingTheory.Polynomial.IntegralNormalization import Mathlib.RingTheory.Polynomial.IrreducibleRing import Mathlib.RingTheory.Polynomial.Nilpotent diff --git a/Mathlib/Algebra/Module/Submodule/Pointwise.lean b/Mathlib/Algebra/Module/Submodule/Pointwise.lean index 99f269138c452..45fc65acdabe9 100644 --- a/Mathlib/Algebra/Module/Submodule/Pointwise.lean +++ b/Mathlib/Algebra/Module/Submodule/Pointwise.lean @@ -3,11 +3,10 @@ Copyright (c) 2021 Eric Wieser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser, Jujian Zhang -/ -import Mathlib.Algebra.Module.BigOperators import Mathlib.Algebra.Group.Subgroup.Pointwise import Mathlib.Algebra.Order.Group.Action -import Mathlib.RingTheory.Ideal.Span -import Mathlib.LinearAlgebra.Finsupp.LinearCombination +import Mathlib.LinearAlgebra.Finsupp.Supported +import Mathlib.LinearAlgebra.Span.Basic /-! # Pointwise instances on `Submodule`s @@ -39,6 +38,7 @@ Other than section `set_acting_on_submodules`, most of the lemmas in this file a lemmas from the file `Mathlib.Algebra.Group.Submonoid.Pointwise`. -/ +assert_not_exists Ideal variable {α : Type*} {R : Type*} {M : Type*} @@ -526,33 +526,6 @@ lemma sup_set_smul (s t : Set S) : · exact Submodule.mem_sup_right (mem_set_smul_of_mem_mem hr hn)) (sup_le (set_smul_mono_left _ le_sup_left) (set_smul_mono_left _ le_sup_right)) -lemma coe_span_smul {R' M' : Type*} [CommSemiring R'] [AddCommMonoid M'] [Module R' M'] - (s : Set R') (N : Submodule R' M') : - (Ideal.span s : Set R') • N = s • N := - set_smul_eq_of_le _ _ _ - (by rintro r n hr hn - induction hr using Submodule.span_induction with - | mem _ h => exact mem_set_smul_of_mem_mem h hn - | zero => rw [zero_smul]; exact Submodule.zero_mem _ - | add _ _ _ _ ihr ihs => rw [add_smul]; exact Submodule.add_mem _ ihr ihs - | smul _ _ hr => - rw [mem_span_set] at hr - obtain ⟨c, hc, rfl⟩ := hr - rw [Finsupp.sum, Finset.smul_sum, Finset.sum_smul] - refine Submodule.sum_mem _ fun i hi => ?_ - rw [← mul_smul, smul_eq_mul, mul_comm, mul_smul] - exact mem_set_smul_of_mem_mem (hc hi) <| Submodule.smul_mem _ _ hn) <| - set_smul_mono_left _ Submodule.subset_span - end set_acting_on_submodules -lemma span_singleton_toAddSubgroup_eq_zmultiples (a : ℤ) : - (span ℤ {a}).toAddSubgroup = AddSubgroup.zmultiples a := by - ext i - simp [Ideal.mem_span_singleton', AddSubgroup.mem_zmultiples_iff] - end Submodule - -@[simp] lemma Ideal.span_singleton_toAddSubgroup_eq_zmultiples (a : ℤ) : - (Ideal.span {a}).toAddSubgroup = AddSubgroup.zmultiples a := - Submodule.span_singleton_toAddSubgroup_eq_zmultiples _ diff --git a/Mathlib/Algebra/Polynomial/AlgebraMap.lean b/Mathlib/Algebra/Polynomial/AlgebraMap.lean index 3431db816584a..b58d9b56eb9b1 100644 --- a/Mathlib/Algebra/Polynomial/AlgebraMap.lean +++ b/Mathlib/Algebra/Polynomial/AlgebraMap.lean @@ -4,11 +4,12 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Johannes Hölzl, Kim Morrison, Jens Wagemaker -/ import Mathlib.Algebra.Algebra.Pi +import Mathlib.Algebra.Algebra.Subalgebra.Basic +import Mathlib.Algebra.Algebra.Tower import Mathlib.Algebra.MonoidAlgebra.Basic import Mathlib.Algebra.Polynomial.Eval.Algebra import Mathlib.Algebra.Polynomial.Eval.Degree import Mathlib.Algebra.Polynomial.Monomial -import Mathlib.RingTheory.Adjoin.Basic /-! # Theory of univariate polynomials @@ -17,6 +18,7 @@ We show that `A[X]` is an R-algebra when `A` is an R-algebra. We promote `eval₂` to an algebra hom in `aeval`. -/ +assert_not_exists Ideal noncomputable section @@ -230,13 +232,6 @@ This is a stronger variant of the linear map `Polynomial.leval`. -/ def aeval : R[X] →ₐ[R] A := eval₂AlgHom' (Algebra.ofId _ _) x (Algebra.commutes · _) -@[simp] -theorem adjoin_X : Algebra.adjoin R ({X} : Set R[X]) = ⊤ := by - refine top_unique fun p _hp => ?_ - set S := Algebra.adjoin R ({X} : Set R[X]) - rw [← sum_monomial_eq p]; simp only [← smul_X_eq_monomial, Sum] - exact S.sum_mem fun n _hn => S.smul_mem (S.pow_mem (Algebra.subset_adjoin rfl) _) _ - @[ext 1200] theorem algHom_ext {f g : R[X] →ₐ[R] B} (hX : f X = g X) : f = g := @@ -415,31 +410,6 @@ theorem aeval_eq_zero_of_dvd_aeval_eq_zero [CommSemiring S] [CommSemiring T] [Al rw [aeval_def, ← eval_map] at h₂ ⊢ exact eval_eq_zero_of_dvd_of_eval_eq_zero (Polynomial.map_dvd (algebraMap S T) h₁) h₂ -variable (R) - -theorem _root_.Algebra.adjoin_singleton_eq_range_aeval (x : A) : - Algebra.adjoin R {x} = (Polynomial.aeval x).range := by - rw [← Algebra.map_top, ← adjoin_X, AlgHom.map_adjoin, Set.image_singleton, aeval_X] - -@[simp] -theorem aeval_mem_adjoin_singleton : - aeval x p ∈ Algebra.adjoin R {x} := by - simpa only [Algebra.adjoin_singleton_eq_range_aeval] using Set.mem_range_self p - -instance instCommSemiringAdjoinSingleton : - CommSemiring <| Algebra.adjoin R {x} := - { mul_comm := fun ⟨p, hp⟩ ⟨q, hq⟩ ↦ by - obtain ⟨p', rfl⟩ := Algebra.adjoin_singleton_eq_range_aeval R x ▸ hp - obtain ⟨q', rfl⟩ := Algebra.adjoin_singleton_eq_range_aeval R x ▸ hq - simp only [AlgHom.toRingHom_eq_coe, RingHom.coe_coe, MulMemClass.mk_mul_mk, ← map_mul, - mul_comm p' q'] } - -instance instCommRingAdjoinSingleton {R A : Type*} [CommRing R] [Ring A] [Algebra R A] (x : A) : - CommRing <| Algebra.adjoin R {x} := - { mul_comm := mul_comm } - -variable {R} - section Semiring variable [Semiring S] {f : R →+* S} diff --git a/Mathlib/Algebra/Polynomial/Derivation.lean b/Mathlib/Algebra/Polynomial/Derivation.lean index 5a14afb6cc51f..b069960426f67 100644 --- a/Mathlib/Algebra/Polynomial/Derivation.lean +++ b/Mathlib/Algebra/Polynomial/Derivation.lean @@ -6,6 +6,7 @@ Authors: Kevin Buzzard, Richard M. Hill import Mathlib.Algebra.Polynomial.AlgebraMap import Mathlib.Algebra.Polynomial.Derivative import Mathlib.Algebra.Polynomial.Module.AEval +import Mathlib.RingTheory.Adjoin.Polynomial import Mathlib.RingTheory.Derivation.Basic /-! # Derivations of univariate polynomials diff --git a/Mathlib/Algebra/Polynomial/Div.lean b/Mathlib/Algebra/Polynomial/Div.lean index 03aa336587516..42daf78354251 100644 --- a/Mathlib/Algebra/Polynomial/Div.lean +++ b/Mathlib/Algebra/Polynomial/Div.lean @@ -3,10 +3,11 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Johannes Hölzl, Kim Morrison, Jens Wagemaker -/ +import Mathlib.Algebra.Field.IsField import Mathlib.Algebra.Polynomial.Inductions import Mathlib.Algebra.Polynomial.Monic +import Mathlib.Algebra.Ring.Regular import Mathlib.RingTheory.Multiplicity -import Mathlib.RingTheory.Ideal.Maps /-! # Division of univariate polynomials @@ -16,7 +17,6 @@ The compatibility between these is given by `modByMonic_add_div`. We also define `rootMultiplicity`. -/ - noncomputable section open Polynomial @@ -592,16 +592,6 @@ theorem dvd_iff_isRoot : X - C a ∣ p ↔ IsRoot p a := theorem X_sub_C_dvd_sub_C_eval : X - C a ∣ p - C (p.eval a) := by rw [dvd_iff_isRoot, IsRoot, eval_sub, eval_C, sub_self] -theorem mem_span_C_X_sub_C_X_sub_C_iff_eval_eval_eq_zero {b : R[X]} {P : R[X][X]} : - P ∈ Ideal.span {C (X - C a), X - C b} ↔ (P.eval b).eval a = 0 := by - rw [Ideal.mem_span_pair] - constructor <;> intro h - · rcases h with ⟨_, _, rfl⟩ - simp only [eval_C, eval_X, eval_add, eval_sub, eval_mul, add_zero, mul_zero, sub_self] - · rcases dvd_iff_isRoot.mpr h with ⟨p, hp⟩ - rcases @X_sub_C_dvd_sub_C_eval _ b _ P with ⟨q, hq⟩ - exact ⟨C p, q, by rw [mul_comm, mul_comm q, eq_add_of_sub_eq' hq, hp, C_mul]⟩ - -- TODO: generalize this to Ring. In general, 0 can be replaced by any element in the center of R. theorem modByMonic_X (p : R[X]) : p %ₘ X = C (p.eval 0) := by rw [← modByMonic_X_sub_C_eq_C_eval, C_0, sub_zero] @@ -615,10 +605,6 @@ theorem sub_dvd_eval_sub (a b : R) (p : R[X]) : a - b ∣ p.eval a - p.eval b := simpa only [coe_evalRingHom, eval_sub, eval_X, eval_C] using (evalRingHom a).map_dvd this simp [dvd_iff_isRoot] -theorem ker_evalRingHom (x : R) : RingHom.ker (evalRingHom x) = Ideal.span {X - C x} := by - ext y - simp [Ideal.mem_span_singleton, dvd_iff_isRoot, RingHom.mem_ker] - @[simp] theorem rootMultiplicity_eq_zero_iff {p : R[X]} {x : R} : rootMultiplicity x p = 0 ↔ IsRoot p x → p = 0 := by diff --git a/Mathlib/Algebra/Polynomial/PartialFractions.lean b/Mathlib/Algebra/Polynomial/PartialFractions.lean index e9bb4f3a58e49..35ee211386718 100644 --- a/Mathlib/Algebra/Polynomial/PartialFractions.lean +++ b/Mathlib/Algebra/Polynomial/PartialFractions.lean @@ -5,6 +5,7 @@ Authors: Kevin Buzzard, Sidharth Hariharan -/ import Mathlib.Algebra.Polynomial.Div import Mathlib.Logic.Function.Basic +import Mathlib.RingTheory.Coprime.Lemmas import Mathlib.RingTheory.Localization.FractionRing import Mathlib.Tactic.FieldSimp import Mathlib.Tactic.LinearCombination diff --git a/Mathlib/Algebra/Polynomial/RingDivision.lean b/Mathlib/Algebra/Polynomial/RingDivision.lean index ca3c0ccfff641..b8bd2125c5b79 100644 --- a/Mathlib/Algebra/Polynomial/RingDivision.lean +++ b/Mathlib/Algebra/Polynomial/RingDivision.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Johannes Hölzl, Kim Morrison, Jens Wagemaker, Johan Commelin -/ import Mathlib.Algebra.Polynomial.AlgebraMap -import Mathlib.Algebra.Polynomial.Degree.Lemmas import Mathlib.Algebra.Polynomial.Div +import Mathlib.RingTheory.Coprime.Basic /-! # Theory of univariate polynomials @@ -14,6 +14,8 @@ We prove basic results about univariate polynomials. -/ +assert_not_exists Ideal.map + noncomputable section open Polynomial @@ -65,11 +67,6 @@ theorem mem_ker_modByMonic (hq : q.Monic) {p : R[X]} : p ∈ LinearMap.ker (modByMonicHom q) ↔ q ∣ p := LinearMap.mem_ker.trans (modByMonic_eq_zero_iff_dvd hq) -@[simp] -theorem ker_modByMonicHom (hq : q.Monic) : - LinearMap.ker (Polynomial.modByMonicHom q) = (Ideal.span {q}).restrictScalars R := - Submodule.ext fun _ => (mem_ker_modByMonic hq).trans Ideal.mem_span_singleton.symm - section variable [Ring S] diff --git a/Mathlib/Algebra/Polynomial/Roots.lean b/Mathlib/Algebra/Polynomial/Roots.lean index 758a417dfb770..12df720582d0e 100644 --- a/Mathlib/Algebra/Polynomial/Roots.lean +++ b/Mathlib/Algebra/Polynomial/Roots.lean @@ -5,8 +5,8 @@ Authors: Chris Hughes, Johannes Hölzl, Kim Morrison, Jens Wagemaker, Johan Comm -/ import Mathlib.Algebra.Polynomial.BigOperators import Mathlib.Algebra.Polynomial.RingDivision -import Mathlib.Data.Fintype.Pi import Mathlib.Data.Set.Finite.Lemmas +import Mathlib.RingTheory.Coprime.Lemmas import Mathlib.RingTheory.Localization.FractionRing import Mathlib.SetTheory.Cardinal.Basic @@ -29,6 +29,8 @@ We define the multiset of roots of a polynomial, and prove basic results about i -/ +assert_not_exists Ideal + open Multiset Finset noncomputable section diff --git a/Mathlib/Algebra/Polynomial/Splits.lean b/Mathlib/Algebra/Polynomial/Splits.lean index 19552f5e13bb4..eb94c89f217df 100644 --- a/Mathlib/Algebra/Polynomial/Splits.lean +++ b/Mathlib/Algebra/Polynomial/Splits.lean @@ -6,6 +6,7 @@ Authors: Chris Hughes import Mathlib.Algebra.Polynomial.FieldDivision import Mathlib.Algebra.Polynomial.Lifts import Mathlib.Data.List.Prime +import Mathlib.RingTheory.Adjoin.Basic import Mathlib.RingTheory.Polynomial.Tower /-! diff --git a/Mathlib/Analysis/Normed/Field/Basic.lean b/Mathlib/Analysis/Normed/Field/Basic.lean index 00424d49f8973..d730d9d8baf82 100644 --- a/Mathlib/Analysis/Normed/Field/Basic.lean +++ b/Mathlib/Analysis/Normed/Field/Basic.lean @@ -9,6 +9,7 @@ import Mathlib.Algebra.Field.Subfield.Defs import Mathlib.Algebra.Order.Group.Pointwise.Interval import Mathlib.Analysis.Normed.Group.Constructions import Mathlib.Analysis.Normed.Group.Submodule +import Mathlib.Algebra.Ring.Regular /-! # Normed fields diff --git a/Mathlib/FieldTheory/IntermediateField/Basic.lean b/Mathlib/FieldTheory/IntermediateField/Basic.lean index 9a52885edff92..0e047a0ddb3f5 100644 --- a/Mathlib/FieldTheory/IntermediateField/Basic.lean +++ b/Mathlib/FieldTheory/IntermediateField/Basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Anne Baanen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ +import Mathlib.Algebra.Algebra.Subalgebra.Tower import Mathlib.Algebra.Field.IsField import Mathlib.Algebra.Field.Subfield.Basic import Mathlib.Algebra.Polynomial.AlgebraMap diff --git a/Mathlib/LinearAlgebra/Matrix/ToLin.lean b/Mathlib/LinearAlgebra/Matrix/ToLin.lean index 16a56b2c45728..de85599206032 100644 --- a/Mathlib/LinearAlgebra/Matrix/ToLin.lean +++ b/Mathlib/LinearAlgebra/Matrix/ToLin.lean @@ -9,6 +9,7 @@ import Mathlib.Data.Matrix.Block import Mathlib.Data.Matrix.Notation import Mathlib.LinearAlgebra.Matrix.StdBasis import Mathlib.RingTheory.AlgebraTower +import Mathlib.RingTheory.Ideal.Span /-! # Linear maps and matrices diff --git a/Mathlib/RingTheory/Adjoin/Basic.lean b/Mathlib/RingTheory/Adjoin/Basic.lean index 17423ca350195..7d4037bf5b2d6 100644 --- a/Mathlib/RingTheory/Adjoin/Basic.lean +++ b/Mathlib/RingTheory/Adjoin/Basic.lean @@ -22,6 +22,7 @@ adjoin, algebra -/ +assert_not_exists Polynomial universe uR uS uA uB diff --git a/Mathlib/RingTheory/Adjoin/Polynomial.lean b/Mathlib/RingTheory/Adjoin/Polynomial.lean new file mode 100644 index 0000000000000..9a8e63128ef2b --- /dev/null +++ b/Mathlib/RingTheory/Adjoin/Polynomial.lean @@ -0,0 +1,67 @@ +/- +Copyright (c) 2018 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes, Johannes Hölzl, Kim Morrison, Jens Wagemaker +-/ +import Mathlib.Algebra.Polynomial.AlgebraMap +import Mathlib.RingTheory.Adjoin.Basic + +/-! +# Polynomials and adjoining roots + +## Main results + +* `Polynomial.instCommSemiringAdjoinSingleton, instCommRingAdjoinSingleton`: + adjoining an element to a commutative (semi)ring gives a commutative (semi)ring +-/ + +noncomputable section + +open Finset + +open Polynomial + +namespace Polynomial + +universe u v w z + +variable {R : Type u} {S : Type v} {T : Type w} {A : Type z} {A' B : Type*} {a b : R} {n : ℕ} + +section aeval + +variable [CommSemiring R] [Semiring A] [CommSemiring A'] [Semiring B] +variable [Algebra R A] [Algebra R B] +variable {p q : R[X]} (x : A) + +@[simp] +theorem adjoin_X : Algebra.adjoin R ({X} : Set R[X]) = ⊤ := by + refine top_unique fun p _hp => ?_ + set S := Algebra.adjoin R ({X} : Set R[X]) + rw [← sum_monomial_eq p]; simp only [← smul_X_eq_monomial, Sum] + exact S.sum_mem fun n _hn => S.smul_mem (S.pow_mem (Algebra.subset_adjoin rfl) _) _ + +variable (R) +theorem _root_.Algebra.adjoin_singleton_eq_range_aeval (x : A) : + Algebra.adjoin R {x} = (Polynomial.aeval x).range := by + rw [← Algebra.map_top, ← adjoin_X, AlgHom.map_adjoin, Set.image_singleton, aeval_X] + +@[simp] +theorem aeval_mem_adjoin_singleton : + aeval x p ∈ Algebra.adjoin R {x} := by + simpa only [Algebra.adjoin_singleton_eq_range_aeval] using Set.mem_range_self p + +instance instCommSemiringAdjoinSingleton : + CommSemiring <| Algebra.adjoin R {x} := + { mul_comm := fun ⟨p, hp⟩ ⟨q, hq⟩ ↦ by + obtain ⟨p', rfl⟩ := Algebra.adjoin_singleton_eq_range_aeval R x ▸ hp + obtain ⟨q', rfl⟩ := Algebra.adjoin_singleton_eq_range_aeval R x ▸ hq + simp only [AlgHom.toRingHom_eq_coe, RingHom.coe_coe, MulMemClass.mk_mul_mk, ← map_mul, + mul_comm p' q'] } + +instance instCommRingAdjoinSingleton {R A : Type*} [CommRing R] [Ring A] [Algebra R A] (x : A) : + CommRing <| Algebra.adjoin R {x} := + { mul_comm := mul_comm } + +end aeval + +end Polynomial diff --git a/Mathlib/RingTheory/Algebraic/MvPolynomial.lean b/Mathlib/RingTheory/Algebraic/MvPolynomial.lean index 649cec44038ad..8ff2e80a728e9 100644 --- a/Mathlib/RingTheory/Algebraic/MvPolynomial.lean +++ b/Mathlib/RingTheory/Algebraic/MvPolynomial.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ import Mathlib.Algebra.MvPolynomial.Supported +import Mathlib.RingTheory.Adjoin.Polynomial import Mathlib.RingTheory.Algebraic.Basic /-! diff --git a/Mathlib/RingTheory/FiniteType.lean b/Mathlib/RingTheory/FiniteType.lean index e8cd7c64fe86f..40f8bd11de976 100644 --- a/Mathlib/RingTheory/FiniteType.lean +++ b/Mathlib/RingTheory/FiniteType.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ import Mathlib.Algebra.FreeAlgebra +import Mathlib.RingTheory.Adjoin.Polynomial import Mathlib.RingTheory.Adjoin.Tower import Mathlib.RingTheory.Ideal.Quotient.Operations import Mathlib.RingTheory.Noetherian.Orzech diff --git a/Mathlib/RingTheory/HahnSeries/Multiplication.lean b/Mathlib/RingTheory/HahnSeries/Multiplication.lean index 7719c9ebf8a30..8660ddb736c28 100644 --- a/Mathlib/RingTheory/HahnSeries/Multiplication.lean +++ b/Mathlib/RingTheory/HahnSeries/Multiplication.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson, Scott Carnahan -/ import Mathlib.Algebra.Algebra.Subalgebra.Basic +import Mathlib.Algebra.Module.BigOperators import Mathlib.Data.Finset.MulAntidiagonal import Mathlib.Data.Finset.SMulAntidiagonal import Mathlib.GroupTheory.GroupAction.Ring diff --git a/Mathlib/RingTheory/Ideal/Operations.lean b/Mathlib/RingTheory/Ideal/Operations.lean index af540403ae860..d627344ccb869 100644 --- a/Mathlib/RingTheory/Ideal/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Operations.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ import Mathlib.Algebra.Algebra.Operations +import Mathlib.Algebra.Module.BigOperators import Mathlib.Data.Fintype.Lattice import Mathlib.RingTheory.Coprime.Lemmas import Mathlib.RingTheory.Ideal.Basic @@ -22,6 +23,33 @@ open Pointwise namespace Submodule +lemma coe_span_smul {R' M' : Type*} [CommSemiring R'] [AddCommMonoid M'] [Module R' M'] + (s : Set R') (N : Submodule R' M') : + (Ideal.span s : Set R') • N = s • N := + set_smul_eq_of_le _ _ _ + (by rintro r n hr hn + induction hr using Submodule.span_induction with + | mem _ h => exact mem_set_smul_of_mem_mem h hn + | zero => rw [zero_smul]; exact Submodule.zero_mem _ + | add _ _ _ _ ihr ihs => rw [add_smul]; exact Submodule.add_mem _ ihr ihs + | smul _ _ hr => + rw [mem_span_set] at hr + obtain ⟨c, hc, rfl⟩ := hr + rw [Finsupp.sum, Finset.smul_sum, Finset.sum_smul] + refine Submodule.sum_mem _ fun i hi => ?_ + rw [← mul_smul, smul_eq_mul, mul_comm, mul_smul] + exact mem_set_smul_of_mem_mem (hc hi) <| Submodule.smul_mem _ _ hn) <| + set_smul_mono_left _ Submodule.subset_span + +lemma span_singleton_toAddSubgroup_eq_zmultiples (a : ℤ) : + (span ℤ {a}).toAddSubgroup = AddSubgroup.zmultiples a := by + ext i + simp [Ideal.mem_span_singleton', AddSubgroup.mem_zmultiples_iff] + +@[simp] lemma _root_.Ideal.span_singleton_toAddSubgroup_eq_zmultiples (a : ℤ) : + (Ideal.span {a}).toAddSubgroup = AddSubgroup.zmultiples a := + Submodule.span_singleton_toAddSubgroup_eq_zmultiples _ + variable {R : Type u} {M : Type v} {M' F G : Type*} section Semiring diff --git a/Mathlib/RingTheory/IntegralClosure/IsIntegral/Basic.lean b/Mathlib/RingTheory/IntegralClosure/IsIntegral/Basic.lean index c90e0cd0abbb5..a8b4fdf450893 100644 --- a/Mathlib/RingTheory/IntegralClosure/IsIntegral/Basic.lean +++ b/Mathlib/RingTheory/IntegralClosure/IsIntegral/Basic.lean @@ -5,6 +5,7 @@ Authors: Kenny Lau -/ import Mathlib.RingTheory.IntegralClosure.IsIntegral.Defs import Mathlib.Algebra.Polynomial.Expand +import Mathlib.RingTheory.Adjoin.Polynomial import Mathlib.RingTheory.Finiteness.Subalgebra import Mathlib.RingTheory.Polynomial.Tower diff --git a/Mathlib/RingTheory/IsTensorProduct.lean b/Mathlib/RingTheory/IsTensorProduct.lean index 292d28691d254..bbe4d72567496 100644 --- a/Mathlib/RingTheory/IsTensorProduct.lean +++ b/Mathlib/RingTheory/IsTensorProduct.lean @@ -3,8 +3,9 @@ Copyright (c) 2022 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ -import Mathlib.RingTheory.TensorProduct.Basic import Mathlib.Algebra.Module.ULift +import Mathlib.RingTheory.TensorProduct.Basic +import Mathlib.Tactic.Ring /-! # The characteristic predicate of tensor product diff --git a/Mathlib/RingTheory/Localization/Basic.lean b/Mathlib/RingTheory/Localization/Basic.lean index 9159d55e9cdbd..d2ed2970d4c2b 100644 --- a/Mathlib/RingTheory/Localization/Basic.lean +++ b/Mathlib/RingTheory/Localization/Basic.lean @@ -7,7 +7,6 @@ import Mathlib.Algebra.Algebra.Tower import Mathlib.Algebra.Field.IsField import Mathlib.Algebra.GroupWithZero.NonZeroDivisors import Mathlib.GroupTheory.MonoidLocalization.MonoidWithZero -import Mathlib.RingTheory.Ideal.Defs import Mathlib.RingTheory.Localization.Defs import Mathlib.RingTheory.OreLocalization.Ring @@ -68,6 +67,7 @@ localization, ring localization, commutative ring localization, characteristic p commutative ring, field of fractions -/ +assert_not_exists Ideal open Function @@ -82,15 +82,6 @@ section IsLocalization variable [IsLocalization M S] -theorem mk'_mem_iff {x} {y : M} {I : Ideal S} : mk' S x y ∈ I ↔ algebraMap R S x ∈ I := by - constructor <;> intro h - · rw [← mk'_spec S x y, mul_comm] - exact I.mul_mem_left ((algebraMap R S) y) h - · rw [← mk'_spec S x y] at h - obtain ⟨b, hb⟩ := isUnit_iff_exists_inv.1 (map_units S y) - have := I.mul_mem_left b h - rwa [mul_comm, mul_assoc, hb, mul_one] at this - variable (M S) in include M in theorem linearMap_compatibleSMul (N₁ N₂) [AddCommMonoid N₁] [AddCommMonoid N₂] [Module R N₁] diff --git a/Mathlib/RingTheory/Localization/FractionRing.lean b/Mathlib/RingTheory/Localization/FractionRing.lean index 155c5ec2966a3..137b20d13c55e 100644 --- a/Mathlib/RingTheory/Localization/FractionRing.lean +++ b/Mathlib/RingTheory/Localization/FractionRing.lean @@ -33,6 +33,7 @@ localization, ring localization, commutative ring localization, characteristic p commutative ring, field of fractions -/ +assert_not_exists Ideal variable (R : Type*) [CommRing R] {M : Submonoid R} (S : Type*) [CommRing S] variable [Algebra R S] {P : Type*} [CommRing P] diff --git a/Mathlib/RingTheory/Localization/Ideal.lean b/Mathlib/RingTheory/Localization/Ideal.lean index 126d2cbf0c729..fb2341191e944 100644 --- a/Mathlib/RingTheory/Localization/Ideal.lean +++ b/Mathlib/RingTheory/Localization/Ideal.lean @@ -5,7 +5,7 @@ Authors: Kenny Lau, Mario Carneiro, Johan Commelin, Amelia Livingston, Anne Baan -/ import Mathlib.GroupTheory.MonoidLocalization.Away import Mathlib.RingTheory.Ideal.Quotient.Operations -import Mathlib.RingTheory.Localization.Basic +import Mathlib.RingTheory.Localization.Defs /-! # Ideals in localizations of commutative rings @@ -24,6 +24,16 @@ section CommSemiring variable {R : Type*} [CommSemiring R] (M : Submonoid R) (S : Type*) [CommSemiring S] variable [Algebra R S] [IsLocalization M S] +variable {M S} in +theorem mk'_mem_iff {x} {y : M} {I : Ideal S} : mk' S x y ∈ I ↔ algebraMap R S x ∈ I := by + constructor <;> intro h + · rw [← mk'_spec S x y, mul_comm] + exact I.mul_mem_left ((algebraMap R S) y) h + · rw [← mk'_spec S x y] at h + obtain ⟨b, hb⟩ := isUnit_iff_exists_inv.1 (map_units S y) + have := I.mul_mem_left b h + rwa [mul_comm, mul_assoc, hb, mul_one] at this + /-- Explicit characterization of the ideal given by `Ideal.map (algebraMap R S) I`. In practice, this ideal differs only in that the carrier set is defined explicitly. This definition is only meant to be used in proving `mem_map_algebraMap_iff`, diff --git a/Mathlib/RingTheory/Polynomial/Basic.lean b/Mathlib/RingTheory/Polynomial/Basic.lean index 4bee44c7de8c5..5513563c4b0da 100644 --- a/Mathlib/RingTheory/Polynomial/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Basic.lean @@ -750,37 +750,6 @@ theorem is_fg_degreeLE [IsNoetherianRing R] (I : Ideal R[X]) (n : ℕ) : isNoetherian_submodule_left.1 (isNoetherian_of_fg_of_noetherian _ ⟨_, degreeLE_eq_span_X_pow.symm⟩) _ -open Algebra in -lemma _root_.Algebra.mem_ideal_map_adjoin {R S : Type*} [CommRing R] [CommRing S] [Algebra R S] - (x : S) (I : Ideal R) {y : adjoin R ({x} : Set S)} : - y ∈ I.map (algebraMap R (adjoin R ({x} : Set S))) ↔ - ∃ p : R[X], (∀ i, p.coeff i ∈ I) ∧ Polynomial.aeval x p = y := by - constructor - · intro H - induction' H using Submodule.span_induction with a ha a b ha hb ha' hb' a b hb hb' - · obtain ⟨a, ha, rfl⟩ := ha - exact ⟨C a, fun i ↦ by rw [coeff_C]; aesop, aeval_C _ _⟩ - · exact ⟨0, by simp, aeval_zero _⟩ - · obtain ⟨a, ha, ha'⟩ := ha' - obtain ⟨b, hb, hb'⟩ := hb' - exact ⟨a + b, fun i ↦ by simpa using add_mem (ha i) (hb i), by simp [ha', hb']⟩ - · obtain ⟨b', hb, hb'⟩ := hb' - obtain ⟨a, ha⟩ := a - rw [Algebra.adjoin_singleton_eq_range_aeval] at ha - obtain ⟨p, hp : aeval x p = a⟩ := ha - refine ⟨p * b', fun i ↦ ?_, by simp [hp, hb']⟩ - rw [coeff_mul] - exact sum_mem fun i hi ↦ Ideal.mul_mem_left _ _ (hb _) - · rintro ⟨p, hp, hp'⟩ - have : y = ∑ i in p.support, p.coeff i • ⟨_, (X ^ i).aeval_mem_adjoin_singleton _ x⟩ := by - trans ∑ i in p.support, ⟨_, (C (p.coeff i) * X ^ i).aeval_mem_adjoin_singleton _ x⟩ - · ext1 - simp only [AddSubmonoidClass.coe_finset_sum, ← map_sum, ← hp', ← as_sum_support_C_mul_X_pow] - · congr with i - simp [Algebra.smul_def] - simp_rw [this, Algebra.smul_def] - exact sum_mem fun i _ ↦ Ideal.mul_mem_right _ _ (Ideal.mem_map_of_mem _ (hp i)) - end CommRing end Ideal diff --git a/Mathlib/RingTheory/Polynomial/Eisenstein/Basic.lean b/Mathlib/RingTheory/Polynomial/Eisenstein/Basic.lean index 5e2bf7367a632..a1e63fad9a328 100644 --- a/Mathlib/RingTheory/Polynomial/Eisenstein/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Eisenstein/Basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Riccardo Brasca. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Riccardo Brasca -/ +import Mathlib.RingTheory.Adjoin.Basic import Mathlib.RingTheory.EisensteinCriterion import Mathlib.RingTheory.Polynomial.ScaleRoots diff --git a/Mathlib/RingTheory/Polynomial/Ideal.lean b/Mathlib/RingTheory/Polynomial/Ideal.lean new file mode 100644 index 0000000000000..e6fd7043addce --- /dev/null +++ b/Mathlib/RingTheory/Polynomial/Ideal.lean @@ -0,0 +1,76 @@ +/- +Copyright (c) 2019 Kenny Lau. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kenny Lau +-/ +import Mathlib.Algebra.Polynomial.RingDivision +import Mathlib.RingTheory.Adjoin.Polynomial +import Mathlib.RingTheory.Ideal.Maps + +/-! +# Ideals in polynomial rings +-/ + +noncomputable section + +open Polynomial + +open Finset + +universe u v w + +namespace Polynomial + +variable {R : Type*} [CommRing R] {a : R} + +theorem mem_span_C_X_sub_C_X_sub_C_iff_eval_eval_eq_zero {b : R[X]} {P : R[X][X]} : + P ∈ Ideal.span {C (X - C a), X - C b} ↔ (P.eval b).eval a = 0 := by + rw [Ideal.mem_span_pair] + constructor <;> intro h + · rcases h with ⟨_, _, rfl⟩ + simp only [eval_C, eval_X, eval_add, eval_sub, eval_mul, add_zero, mul_zero, sub_self] + · rcases dvd_iff_isRoot.mpr h with ⟨p, hp⟩ + rcases @X_sub_C_dvd_sub_C_eval _ b _ P with ⟨q, hq⟩ + exact ⟨C p, q, by rw [mul_comm, mul_comm q, eq_add_of_sub_eq' hq, hp, C_mul]⟩ + +theorem ker_evalRingHom (x : R) : RingHom.ker (evalRingHom x) = Ideal.span {X - C x} := by + ext y + simp [Ideal.mem_span_singleton, dvd_iff_isRoot, RingHom.mem_ker] + +@[simp] +theorem ker_modByMonicHom {q : R[X]} (hq : q.Monic) : + LinearMap.ker (Polynomial.modByMonicHom q) = (Ideal.span {q}).restrictScalars R := + Submodule.ext fun _ => (mem_ker_modByMonic hq).trans Ideal.mem_span_singleton.symm + +open Algebra in +lemma _root_.Algebra.mem_ideal_map_adjoin {R S : Type*} [CommRing R] [CommRing S] [Algebra R S] + (x : S) (I : Ideal R) {y : adjoin R ({x} : Set S)} : + y ∈ I.map (algebraMap R (adjoin R ({x} : Set S))) ↔ + ∃ p : R[X], (∀ i, p.coeff i ∈ I) ∧ Polynomial.aeval x p = y := by + constructor + · intro H + induction' H using Submodule.span_induction with a ha a b ha hb ha' hb' a b hb hb' + · obtain ⟨a, ha, rfl⟩ := ha + exact ⟨C a, fun i ↦ by rw [coeff_C]; aesop, aeval_C _ _⟩ + · exact ⟨0, by simp, aeval_zero _⟩ + · obtain ⟨a, ha, ha'⟩ := ha' + obtain ⟨b, hb, hb'⟩ := hb' + exact ⟨a + b, fun i ↦ by simpa using add_mem (ha i) (hb i), by simp [ha', hb']⟩ + · obtain ⟨b', hb, hb'⟩ := hb' + obtain ⟨a, ha⟩ := a + rw [Algebra.adjoin_singleton_eq_range_aeval] at ha + obtain ⟨p, hp : aeval x p = a⟩ := ha + refine ⟨p * b', fun i ↦ ?_, by simp [hp, hb']⟩ + rw [coeff_mul] + exact sum_mem fun i hi ↦ Ideal.mul_mem_left _ _ (hb _) + · rintro ⟨p, hp, hp'⟩ + have : y = ∑ i in p.support, p.coeff i • ⟨_, (X ^ i).aeval_mem_adjoin_singleton _ x⟩ := by + trans ∑ i in p.support, ⟨_, (C (p.coeff i) * X ^ i).aeval_mem_adjoin_singleton _ x⟩ + · ext1 + simp only [AddSubmonoidClass.coe_finset_sum, ← map_sum, ← hp', ← as_sum_support_C_mul_X_pow] + · congr with i + simp [Algebra.smul_def] + simp_rw [this, Algebra.smul_def] + exact sum_mem fun i _ ↦ Ideal.mul_mem_right _ _ (Ideal.mem_map_of_mem _ (hp i)) + +end Polynomial diff --git a/Mathlib/RingTheory/Polynomial/IntegralNormalization.lean b/Mathlib/RingTheory/Polynomial/IntegralNormalization.lean index ee4f04ec120b8..967c8dbcef67d 100644 --- a/Mathlib/RingTheory/Polynomial/IntegralNormalization.lean +++ b/Mathlib/RingTheory/Polynomial/IntegralNormalization.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Johannes Hölzl, Kim Morrison, Jens Wagemaker, Andrew Yang, Yuyang Zhao -/ import Mathlib.Algebra.Polynomial.Monic +import Mathlib.Algebra.Ring.Regular import Mathlib.RingTheory.Polynomial.ScaleRoots /-! diff --git a/Mathlib/RingTheory/Polynomial/Quotient.lean b/Mathlib/RingTheory/Polynomial/Quotient.lean index 96efc3704a697..ec63a3a57ca23 100644 --- a/Mathlib/RingTheory/Polynomial/Quotient.lean +++ b/Mathlib/RingTheory/Polynomial/Quotient.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.Polynomial.Eval.SMul import Mathlib.GroupTheory.GroupAction.Ring import Mathlib.RingTheory.Ideal.Quotient.Operations import Mathlib.RingTheory.Polynomial.Basic +import Mathlib.RingTheory.Polynomial.Ideal /-! # Quotients of polynomial rings From 7b33f44e86f208ccddf0b12824e6796bf669eef7 Mon Sep 17 00:00:00 2001 From: Antoine Chambert-Loir Date: Mon, 2 Dec 2024 11:59:14 +0000 Subject: [PATCH 464/829] feat(RingTheory.MvPowerSeries.Topology): define the product topology on mv power series (#14866) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Let `R` be with `Semiring R` and `TopologicalSpace R` In this file we define the topology on `MvPowerSeries σ R` that corresponds to the simple convergence on its coefficients, aka the product topology. It is the coarsest topology for which all coefficients maps are continuous. When `R` has `UniformSpace R`, we define the corresponding uniform structure. When the type of coefficients has the discrete topology, it corresponds to the topology defined by [bourbaki1981], chapter 4, §4, n°2 Coauthored with María Inés de Frutos Fernández @mariainesdff Co-authored-by: leanprover-community-mathlib4-bot Co-authored-by: mariainesdff --- Mathlib.lean | 2 + Mathlib/Algebra/BigOperators/Ring.lean | 8 + .../RingTheory/MvPowerSeries/PiTopology.lean | 228 ++++++++++++++++++ .../RingTheory/PowerSeries/PiTopology.lean | 205 ++++++++++++++++ 4 files changed, 443 insertions(+) create mode 100644 Mathlib/RingTheory/MvPowerSeries/PiTopology.lean create mode 100644 Mathlib/RingTheory/PowerSeries/PiTopology.lean diff --git a/Mathlib.lean b/Mathlib.lean index fccdcac503167..63a705f4c42ae 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4435,6 +4435,7 @@ import Mathlib.RingTheory.MvPowerSeries.Inverse import Mathlib.RingTheory.MvPowerSeries.LexOrder import Mathlib.RingTheory.MvPowerSeries.NoZeroDivisors import Mathlib.RingTheory.MvPowerSeries.Order +import Mathlib.RingTheory.MvPowerSeries.PiTopology import Mathlib.RingTheory.MvPowerSeries.Trunc import Mathlib.RingTheory.Nakayama import Mathlib.RingTheory.Nilpotent.Basic @@ -4497,6 +4498,7 @@ import Mathlib.RingTheory.PowerSeries.Basic import Mathlib.RingTheory.PowerSeries.Derivative import Mathlib.RingTheory.PowerSeries.Inverse import Mathlib.RingTheory.PowerSeries.Order +import Mathlib.RingTheory.PowerSeries.PiTopology import Mathlib.RingTheory.PowerSeries.Trunc import Mathlib.RingTheory.PowerSeries.WellKnown import Mathlib.RingTheory.Presentation diff --git a/Mathlib/Algebra/BigOperators/Ring.lean b/Mathlib/Algebra/BigOperators/Ring.lean index ff6ed790f07e8..b7468a17eabc1 100644 --- a/Mathlib/Algebra/BigOperators/Ring.lean +++ b/Mathlib/Algebra/BigOperators/Ring.lean @@ -178,6 +178,14 @@ theorem prod_add (f g : ι → α) (s : Finset ι) : simp only [mem_filter, mem_sdiff, not_and, not_exists, and_congr_right_iff] tauto) +theorem prod_one_add {f : ι → α} (s : Finset ι) : + ∏ i ∈ s, (1 + f i) = ∑ t ∈ s.powerset, ∏ i ∈ t, f i := by + simp only [add_comm (1 : α), prod_add, prod_const_one, mul_one] + +theorem prod_add_one {f : ι → α} (s : Finset ι) : + ∏ i ∈ s, (f i + 1) = ∑ t ∈ s.powerset, ∏ i ∈ t, f i := by + simp only [prod_add, prod_const_one, mul_one] + end DecidableEq /-- `∏ i, (f i + g i) = (∏ i, f i) + ∑ i, g i * (∏ j < i, f j + g j) * (∏ j > i, f j)`. -/ diff --git a/Mathlib/RingTheory/MvPowerSeries/PiTopology.lean b/Mathlib/RingTheory/MvPowerSeries/PiTopology.lean new file mode 100644 index 0000000000000..f3335ced2bf63 --- /dev/null +++ b/Mathlib/RingTheory/MvPowerSeries/PiTopology.lean @@ -0,0 +1,228 @@ +/- +Copyright (c) 2024 Antoine Chambert-Loir, María Inés de Frutos Fernández. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Antoine Chambert-Loir, María Inés de Frutos Fernández +-/ + +import Mathlib.RingTheory.Nilpotent.Defs +import Mathlib.RingTheory.MvPowerSeries.Basic +import Mathlib.Topology.Algebra.InfiniteSum.Constructions +import Mathlib.Topology.Algebra.Ring.Basic +import Mathlib.Topology.Algebra.UniformGroup.Basic +import Mathlib.Topology.UniformSpace.Pi + +/-! # Product topology on multivariate power series + +Let `R` be with `Semiring R` and `TopologicalSpace R` +In this file we define the topology on `MvPowerSeries σ R` +that corresponds to the simple convergence on its coefficients. +It is the coarsest topology for which all coefficient maps are continuous. + +When `R` has `UniformSpace R`, we define the corresponding uniform structure. + +This topology can be included by writing `open scoped MvPowerSeries.WithPiTopology`. + +When the type of coefficients has the discrete topology, +it corresponds to the topology defined by [bourbaki1981], chapter 4, §4, n°2. + +It is *not* the adic topology in general. + +## Main results + +- `MvPowerSeries.WithPiTopology.tendsto_pow_zero_of_constantCoeff_nilpotent`, + `MvPowerSeries.WithPiTopology.tendsto_pow_zero_of_constantCoeff_zero`: if the constant coefficient + of `f` is nilpotent, or vanishes, then the powers of `f` converge to zero. + +- `MvPowerSeries.WithPiTopology.tendsto_pow_of_constantCoeff_nilpotent_iff` : the powers of `f` + converge to zero iff the constant coefficient of `f` is nilpotent. + +- `MvPowerSeries.WithPiTopology.hasSum_of_monomials_self` : viewed as an infinite sum, a power + series converges to itself. + +TODO: add the similar result for the series of homogeneous components. + +## Instances + +- If `R` is a topological (semi)ring, then so is `MvPowerSeries σ R`. +- If the topology of `R` is T0 or T2, then so is that of `MvPowerSeries σ R`. +- If `R` is a `UniformAddGroup`, then so is `MvPowerSeries σ R`. +- If `R` is complete, then so is `MvPowerSeries σ R`. + +-/ + +namespace MvPowerSeries + +open Function Filter + +variable {σ R : Type*} + +namespace WithPiTopology + +section Topology + +variable [TopologicalSpace R] + +variable (R) in +/-- The pointwise topology on `MvPowerSeries` -/ +scoped instance : TopologicalSpace (MvPowerSeries σ R) := + Pi.topologicalSpace + +/-- `MvPowerSeries` on a `T0Space` form a `T0Space` -/ +@[scoped instance] +theorem instT0Space [T0Space R] : T0Space (MvPowerSeries σ R) := Pi.instT0Space + +/-- `MvPowerSeries` on a `T2Space` form a `T2Space` -/ +@[scoped instance] +theorem instT2Space [T2Space R] : T2Space (MvPowerSeries σ R) := Pi.t2Space + +variable (R) in +/-- `MvPowerSeries.coeff` is continuous. -/ +@[fun_prop] +theorem continuous_coeff [Semiring R] (d : σ →₀ ℕ) : + Continuous (MvPowerSeries.coeff R d) := + continuous_pi_iff.mp continuous_id d + +variable (R) in +/-- `MvPolynomial.constantCoeff` is continuous -/ +theorem continuous_constantCoeff [Semiring R] : Continuous (constantCoeff σ R) := + continuous_coeff R 0 + +/-- A family of power series converges iff it converges coefficientwise -/ +theorem tendsto_iff_coeff_tendsto [Semiring R] {ι : Type*} + (f : ι → MvPowerSeries σ R) (u : Filter ι) (g : MvPowerSeries σ R) : + Tendsto f u (nhds g) ↔ + ∀ d : σ →₀ ℕ, Tendsto (fun i => coeff R d (f i)) u (nhds (coeff R d g)) := by + rw [nhds_pi, tendsto_pi] + exact forall_congr' (fun d => Iff.rfl) + +variable (σ R) + +/-- The semiring topology on `MvPowerSeries` of a topological semiring -/ +@[scoped instance] +theorem instTopologicalSemiring [Semiring R] [TopologicalSemiring R] : + TopologicalSemiring (MvPowerSeries σ R) where + continuous_add := continuous_pi fun d => continuous_add.comp + (((continuous_coeff R d).fst').prod_mk (continuous_coeff R d).snd') + continuous_mul := continuous_pi fun _ => + continuous_finset_sum _ fun i _ => continuous_mul.comp + ((continuous_coeff R i.fst).fst'.prod_mk (continuous_coeff R i.snd).snd') + +/-- The ring topology on `MvPowerSeries` of a topological ring -/ +@[scoped instance] +theorem instTopologicalRing [Ring R] [TopologicalRing R] : + TopologicalRing (MvPowerSeries σ R) := + { instTopologicalSemiring σ R with + continuous_neg := continuous_pi fun d ↦ Continuous.comp continuous_neg + (continuous_coeff R d) } + +variable {σ R} + +@[fun_prop] +theorem continuous_C [Semiring R] : + Continuous (C σ R) := by + classical + simp only [continuous_iff_continuousAt] + refine fun r ↦ (tendsto_iff_coeff_tendsto _ _ _).mpr fun d ↦ ?_ + simp only [coeff_C] + split_ifs + · exact tendsto_id + · exact tendsto_const_nhds + +theorem variables_tendsto_zero [Semiring R] : + Tendsto (X · : σ → MvPowerSeries σ R) cofinite (nhds 0) := by + classical + simp only [tendsto_iff_coeff_tendsto, ← coeff_apply, coeff_X, coeff_zero] + refine fun d ↦ tendsto_nhds_of_eventually_eq ?_ + by_cases h : ∃ i, d = Finsupp.single i 1 + · obtain ⟨i, hi⟩ := h + filter_upwards [eventually_cofinite_ne i] with j hj + simp [hi, Finsupp.single_eq_single_iff, hj.symm] + · simpa only [ite_eq_right_iff] using + Eventually.of_forall fun x h' ↦ (not_exists.mp h x h').elim + +theorem tendsto_pow_zero_of_constantCoeff_nilpotent [CommSemiring R] + {f} (hf : IsNilpotent (constantCoeff σ R f)) : + Tendsto (fun n : ℕ => f ^ n) atTop (nhds 0) := by + classical + obtain ⟨m, hm⟩ := hf + simp_rw [tendsto_iff_coeff_tendsto, coeff_zero] + exact fun d ↦ tendsto_atTop_of_eventually_const fun n hn ↦ + coeff_eq_zero_of_constantCoeff_nilpotent hm hn + +theorem tendsto_pow_zero_of_constantCoeff_zero [CommSemiring R] + {f} (hf : constantCoeff σ R f = 0) : + Tendsto (fun n : ℕ => f ^ n) atTop (nhds 0) := by + apply tendsto_pow_zero_of_constantCoeff_nilpotent + rw [hf] + exact IsNilpotent.zero + +/-- The powers of a `MvPowerSeries` converge to 0 iff its constant coefficient is nilpotent. +N. Bourbaki, *Algebra II*, [bourbaki1981] (chap. 4, §4, n°2, corollaire de la prop. 3) -/ +theorem tendsto_pow_of_constantCoeff_nilpotent_iff [CommRing R] [DiscreteTopology R] (f) : + Tendsto (fun n : ℕ => f ^ n) atTop (nhds 0) ↔ + IsNilpotent (constantCoeff σ R f) := by + refine ⟨?_, tendsto_pow_zero_of_constantCoeff_nilpotent⟩ + intro h + suffices Tendsto (fun n : ℕ => constantCoeff σ R (f ^ n)) atTop (nhds 0) by + simp only [tendsto_def] at this + specialize this {0} _ + suffices ∀ x : R, {x} ∈ nhds x by exact this 0 + rw [← discreteTopology_iff_singleton_mem_nhds]; infer_instance + simp only [map_pow, mem_atTop_sets, ge_iff_le, Set.mem_preimage, + Set.mem_singleton_iff] at this + obtain ⟨m, hm⟩ := this + use m + apply hm m (le_refl m) + simp only [← @comp_apply _ R ℕ, ← tendsto_map'_iff] + simp only [Tendsto, map_le_iff_le_comap] at h ⊢ + refine le_trans h (comap_mono ?_) + rw [← map_le_iff_le_comap] + exact Continuous.continuousAt (continuous_constantCoeff R) + +variable [Semiring R] + +/-- A multivariate power series is the sum (in the sense of summable families) of its monomials -/ +theorem hasSum_of_monomials_self (f : MvPowerSeries σ R) : + HasSum (fun d : σ →₀ ℕ => monomial R d (coeff R d f)) f := by + rw [Pi.hasSum] + intro d + convert hasSum_single d ?_ using 1 + · exact (coeff_monomial_same d _).symm + · exact fun d' h ↦ coeff_monomial_ne (Ne.symm h) _ + +/-- If the coefficient space is T2, then the multivariate power series is `tsum` of its monomials -/ +theorem as_tsum [T2Space R] (f : MvPowerSeries σ R) : + f = tsum fun d : σ →₀ ℕ => monomial R d (coeff R d f) := + (HasSum.tsum_eq (hasSum_of_monomials_self _)).symm + +end Topology + +section Uniformity + +variable [UniformSpace R] + +/-- The componentwise uniformity on `MvPowerSeries` -/ +scoped instance : UniformSpace (MvPowerSeries σ R) := + Pi.uniformSpace fun _ : σ →₀ ℕ => R + +variable (R) in +/-- Coefficients of a multivariate power series are uniformly continuous -/ +theorem uniformContinuous_coeff [Semiring R] (d : σ →₀ ℕ) : + UniformContinuous fun f : MvPowerSeries σ R => coeff R d f := + uniformContinuous_pi.mp uniformContinuous_id d + +/-- Completeness of the uniform structure on `MvPowerSeries` -/ +@[scoped instance] +theorem instCompleteSpace [CompleteSpace R] : + CompleteSpace (MvPowerSeries σ R) := Pi.complete _ + +/-- The `UniformAddGroup` structure on `MvPowerSeries` of a `UniformAddGroup` -/ +@[scoped instance] +theorem instUniformAddGroup [AddGroup R] [UniformAddGroup R] : + UniformAddGroup (MvPowerSeries σ R) := Pi.instUniformAddGroup + +end Uniformity + +end WithPiTopology + +end MvPowerSeries diff --git a/Mathlib/RingTheory/PowerSeries/PiTopology.lean b/Mathlib/RingTheory/PowerSeries/PiTopology.lean new file mode 100644 index 0000000000000..b3cf011bd7091 --- /dev/null +++ b/Mathlib/RingTheory/PowerSeries/PiTopology.lean @@ -0,0 +1,205 @@ +/- +Copyright (c) 2024 Antoine Chambert-Loir, María Inés de Frutos-Fernández. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Antoine Chambert-Loir, María Inés de Frutos-Fernández +-/ +import Mathlib.RingTheory.MvPowerSeries.PiTopology +import Mathlib.RingTheory.PowerSeries.Basic +import Mathlib.LinearAlgebra.Finsupp.Pi + +/-! # Product topology on power series + +Let `R` be with `Semiring R` and `TopologicalSpace R` +In this file we define the topology on `PowerSeries σ R` +that corresponds to the simple convergence on its coefficients. +It is the coarsest topology for which all coefficients maps are continuous. + +When `R` has `UniformSpace R`, we define the corresponding uniform structure. + +This topology can be included by writing `open scoped PowerSeries.WithPiTopology`. + +When the type of coefficients has the discrete topology, +it corresponds to the topology defined by [bourbaki1981], chapter 4, §4, n°2. + +It corresponds with the adic topology but this is not proved here. + +- `PowerSeries.WithPiTopology.tendsto_pow_zero_of_constantCoeff_nilpotent`, +`PowerSeries.WithPiTopology.tendsto_pow_zero_of_constantCoeff_zero`: if the constant coefficient +of `f` is nilpotent, or vanishes, then the powers of `f` converge to zero. + +- `PowerSeries.WithPiTopology.tendsto_pow_zero_of_constantCoeff_nilpotent_iff` : the powers of `f` +converge to zero iff the constant coefficient of `f` is nilpotent. + +- `PowerSeries.WithPiTopology.hasSum_of_monomials_self` : viewed as an infinite sum, a power +series coverges to itself. + +TODO: add the similar result for the series of homogeneous components. + +## Instances + +- If `R` is a topological (semi)ring, then so is `PowerSeries σ R`. +- If the topology of `R` is T0 or T2, then so is that of `PowerSeries σ R`. +- If `R` is a `UniformAddGroup`, then so is `PowerSeries σ R`. +- If `R` is complete, then so is `PowerSeries σ R`. + +-/ + + +namespace PowerSeries + +open Filter Function + +variable (R : Type*) + +section Topological + +variable [TopologicalSpace R] + +namespace WithPiTopology + +/-- The pointwise topology on `PowerSeries` -/ +scoped instance : TopologicalSpace (PowerSeries R) := + Pi.topologicalSpace + +/-- Separation of the topology on `PowerSeries` -/ +@[scoped instance] +theorem instT0Space [T0Space R] : T0Space (PowerSeries R) := + MvPowerSeries.WithPiTopology.instT0Space + +/-- `PowerSeries` on a `T2Space` form a `T2Space` -/ +@[scoped instance] +theorem instT2Space [T2Space R] : T2Space (PowerSeries R) := + MvPowerSeries.WithPiTopology.instT2Space + +/-- Coefficients are continuous -/ +theorem continuous_coeff [Semiring R] (d : ℕ) : Continuous (PowerSeries.coeff R d) := + continuous_pi_iff.mp continuous_id (Finsupp.single () d) + +/-- The constant coefficient is continuous -/ +theorem continuous_constantCoeff [Semiring R] : Continuous (constantCoeff R) := + coeff_zero_eq_constantCoeff (R := R) ▸ continuous_coeff R 0 + +/-- A family of power series converges iff it converges coefficientwise -/ +theorem tendsto_iff_coeff_tendsto [Semiring R] {ι : Type*} + (f : ι → PowerSeries R) (u : Filter ι) (g : PowerSeries R) : + Tendsto f u (nhds g) ↔ + ∀ d : ℕ, Tendsto (fun i => coeff R d (f i)) u (nhds (coeff R d g)) := by + rw [MvPowerSeries.WithPiTopology.tendsto_iff_coeff_tendsto] + apply (Finsupp.LinearEquiv.finsuppUnique ℕ ℕ Unit).toEquiv.forall_congr + intro d + simp only [LinearEquiv.coe_toEquiv, Finsupp.LinearEquiv.finsuppUnique_apply, + PUnit.default_eq_unit, coeff] + apply iff_of_eq + congr + · ext _; congr; ext; simp + · ext; simp + +/-- The semiring topology on `PowerSeries` of a topological semiring -/ +@[scoped instance] +theorem instTopologicalSemiring [Semiring R] [TopologicalSemiring R] : + TopologicalSemiring (PowerSeries R) := + MvPowerSeries.WithPiTopology.instTopologicalSemiring Unit R + +/-- The ring topology on `PowerSeries` of a topological ring -/ +@[scoped instance] +theorem instTopologicalRing [Ring R] [TopologicalRing R] : + TopologicalRing (PowerSeries R) := + MvPowerSeries.WithPiTopology.instTopologicalRing Unit R + +end WithPiTopology + +end Topological + +section Uniform + +namespace WithPiTopology + +variable [UniformSpace R] + +/-- The product uniformity on `PowerSeries` -/ +scoped instance : UniformSpace (PowerSeries R) := + MvPowerSeries.WithPiTopology.instUniformSpace + +/-- Coefficients are uniformly continuous -/ +theorem uniformContinuous_coeff [Semiring R] (d : ℕ) : + UniformContinuous fun f : PowerSeries R ↦ coeff R d f := + uniformContinuous_pi.mp uniformContinuous_id (Finsupp.single () d) + +/-- Completeness of the uniform structure on `PowerSeries` -/ +@[scoped instance] +theorem instCompleteSpace [CompleteSpace R] : + CompleteSpace (PowerSeries R) := + MvPowerSeries.WithPiTopology.instCompleteSpace + +/-- The `UniformAddGroup` structure on `PowerSeries` of a `UniformAddGroup` -/ +@[scoped instance] +theorem instUniformAddGroup [AddGroup R] [UniformAddGroup R] : + UniformAddGroup (PowerSeries R) := + MvPowerSeries.WithPiTopology.instUniformAddGroup + +end WithPiTopology + +end Uniform + +section + +variable {R} + +variable [TopologicalSpace R] + +namespace WithPiTopology + +open MvPowerSeries.WithPiTopology + +theorem continuous_C [Semiring R] : Continuous (C R) := + MvPowerSeries.WithPiTopology.continuous_C + +theorem tendsto_pow_zero_of_constantCoeff_nilpotent [CommSemiring R] + {f : PowerSeries R} (hf : IsNilpotent (constantCoeff R f)) : + Tendsto (fun n : ℕ => f ^ n) atTop (nhds 0) := + MvPowerSeries.WithPiTopology.tendsto_pow_zero_of_constantCoeff_nilpotent hf + +theorem tendsto_pow_zero_of_constantCoeff_zero [CommSemiring R] + {f : PowerSeries R} (hf : constantCoeff R f = 0) : + Tendsto (fun n : ℕ => f ^ n) atTop (nhds 0) := + MvPowerSeries.WithPiTopology.tendsto_pow_zero_of_constantCoeff_zero hf + +/-- The powers of a `PowerSeries` converge to 0 iff its constant coefficient is nilpotent. +N. Bourbaki, *Algebra II*, [bourbaki1981] (chap. 4, §4, n°2, corollaire de la prop. 3) -/ +theorem tendsto_pow_zero_of_constantCoeff_nilpotent_iff + [CommRing R] [DiscreteTopology R] (f : PowerSeries R) : + Tendsto (fun n : ℕ => f ^ n) atTop (nhds 0) ↔ + IsNilpotent (constantCoeff R f) := + MvPowerSeries.WithPiTopology.tendsto_pow_of_constantCoeff_nilpotent_iff f + +end WithPiTopology + +end + +section Summable + +variable [Semiring R] [TopologicalSpace R] + +open WithPiTopology MvPowerSeries.WithPiTopology + +variable {R} + +-- NOTE : one needs an API to apply `Finsupp.LinearEquiv.finsuppUnique` +/-- A power series is the sum (in the sense of summable families) of its monomials -/ +theorem hasSum_of_monomials_self (f : PowerSeries R) : + HasSum (fun d : ℕ => monomial R d (coeff R d f)) f := by + rw [← (Finsupp.LinearEquiv.finsuppUnique ℕ ℕ Unit).toEquiv.hasSum_iff] + convert MvPowerSeries.WithPiTopology.hasSum_of_monomials_self f + simp only [LinearEquiv.coe_toEquiv, comp_apply, monomial, coeff, + Finsupp.LinearEquiv.finsuppUnique_apply, PUnit.default_eq_unit] + congr + all_goals { ext; simp } + +/-- If the coefficient space is T2, then the power series is `tsum` of its monomials -/ +theorem as_tsum [T2Space R] (f : PowerSeries R) : + f = tsum fun d : ℕ => monomial R d (coeff R d f) := + (HasSum.tsum_eq (hasSum_of_monomials_self f)).symm + +end Summable + +end PowerSeries From 41ab76e5430d268338cb239affaf2e42f7993da0 Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Mon, 2 Dec 2024 12:28:46 +0000 Subject: [PATCH 465/829] feat(Algebra/Algebra/Subalgebra/Basic): add `AlgHom.subalgebraMap` (#19641) which is parallel to `LinearMap.submoduleMap`. --- Mathlib/Algebra/Algebra/Subalgebra/Basic.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean b/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean index bf2e80c9aebd3..20f8a381665c2 100644 --- a/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean @@ -938,6 +938,18 @@ variable (f : A →ₐ[R] B) theorem range_comp_val : (f.comp S.val).range = S.map f := by rw [AlgHom.range_comp, range_val] +/-- An `AlgHom` between two rings restricts to an `AlgHom` from any subalgebra of the +domain onto the image of that subalgebra. -/ +def _root_.AlgHom.subalgebraMap : S →ₐ[R] S.map f := + (f.comp S.val).codRestrict _ fun x ↦ ⟨_, x.2, rfl⟩ + +variable {S} in +@[simp] +theorem _root_.AlgHom.subalgebraMap_coe_apply (x : S) : f.subalgebraMap S x = f x := rfl + +theorem _root_.AlgHom.subalgebraMap_surjective : Function.Surjective (f.subalgebraMap S) := + f.toAddMonoidHom.addSubmonoidMap_surjective S.toAddSubmonoid + variable (hf : Function.Injective f) /-- A subalgebra is isomorphic to its image under an injective `AlgHom` -/ From 4662221be714916048d97aa2dc0157c44eca5a51 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Mon, 2 Dec 2024 12:55:52 +0000 Subject: [PATCH 466/829] chore(RingTheory): split `AlgebraicIndependent.lean` (#19606) As a follow-up to splitting `RingTheory/Algebraic.lean`, also split `AlgebraicIndependent.lean`. This PR creates the following new files: * `AlgebraicIndependent/Defs.lean`: definition of `AlgebraicIndependent` and `IsTranscendenceBasis` * `AlgebraicIndependent/Basic.lean`: basic results on `AlgebraicIndependent` * `AlgebraicIndependent/TranscendenceBasis.lean`: basic results on `TranscendenceBasis` * `AlgebraicIndependent/Adjoin.lean`: relating `AlgebraicIndependent` and `IntermediateField.adjoin` * `AlgebraicIndependent/Transcendental.lean`: relating `AlgebraicIndependent`, `Transcendental` and `Algebraic` * `AlgebraicIndependent/RankAndCardinality.lean`: about the rank/cardinality of `TranscendenceBasis`es Co-authored-by: Anne Baanen --- Mathlib.lean | 7 +- .../IsAlgClosed/Classification.lean | 2 +- Mathlib/FieldTheory/SeparableDegree.lean | 3 +- Mathlib/RingTheory/AlgebraicIndependent.lean | 805 ------------------ .../AlgebraicIndependent/Adjoin.lean | 71 ++ .../AlgebraicIndependent/Basic.lean | 384 +++++++++ .../RingTheory/AlgebraicIndependent/Defs.lean | 155 ++++ .../RankAndCardinality.lean | 93 ++ .../TranscendenceBasis.lean | 129 +++ .../AlgebraicIndependent/Transcendental.lean | 146 ++++ 10 files changed, 987 insertions(+), 808 deletions(-) delete mode 100644 Mathlib/RingTheory/AlgebraicIndependent.lean create mode 100644 Mathlib/RingTheory/AlgebraicIndependent/Adjoin.lean create mode 100644 Mathlib/RingTheory/AlgebraicIndependent/Basic.lean create mode 100644 Mathlib/RingTheory/AlgebraicIndependent/Defs.lean create mode 100644 Mathlib/RingTheory/AlgebraicIndependent/RankAndCardinality.lean create mode 100644 Mathlib/RingTheory/AlgebraicIndependent/TranscendenceBasis.lean create mode 100644 Mathlib/RingTheory/AlgebraicIndependent/Transcendental.lean diff --git a/Mathlib.lean b/Mathlib.lean index 63a705f4c42ae..380263015c62c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4227,7 +4227,12 @@ import Mathlib.RingTheory.Algebraic.Integral import Mathlib.RingTheory.Algebraic.LinearIndependent import Mathlib.RingTheory.Algebraic.MvPolynomial import Mathlib.RingTheory.Algebraic.Pi -import Mathlib.RingTheory.AlgebraicIndependent +import Mathlib.RingTheory.AlgebraicIndependent.Adjoin +import Mathlib.RingTheory.AlgebraicIndependent.Basic +import Mathlib.RingTheory.AlgebraicIndependent.Defs +import Mathlib.RingTheory.AlgebraicIndependent.RankAndCardinality +import Mathlib.RingTheory.AlgebraicIndependent.TranscendenceBasis +import Mathlib.RingTheory.AlgebraicIndependent.Transcendental import Mathlib.RingTheory.Artinian import Mathlib.RingTheory.Bezout import Mathlib.RingTheory.Bialgebra.Basic diff --git a/Mathlib/FieldTheory/IsAlgClosed/Classification.lean b/Mathlib/FieldTheory/IsAlgClosed/Classification.lean index 7cd6d8c7278e5..53e821c1bc818 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/Classification.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/Classification.lean @@ -8,7 +8,7 @@ import Mathlib.Algebra.MvPolynomial.Cardinal import Mathlib.Algebra.Polynomial.Cardinal import Mathlib.FieldTheory.IsAlgClosed.Basic import Mathlib.RingTheory.Algebraic.Cardinality -import Mathlib.RingTheory.AlgebraicIndependent +import Mathlib.RingTheory.AlgebraicIndependent.TranscendenceBasis /-! # Classification of Algebraically closed fields diff --git a/Mathlib/FieldTheory/SeparableDegree.lean b/Mathlib/FieldTheory/SeparableDegree.lean index f284a0edd9ea5..1b85e1f399fbf 100644 --- a/Mathlib/FieldTheory/SeparableDegree.lean +++ b/Mathlib/FieldTheory/SeparableDegree.lean @@ -7,7 +7,8 @@ import Mathlib.FieldTheory.SplittingField.Construction import Mathlib.FieldTheory.IsAlgClosed.AlgebraicClosure import Mathlib.FieldTheory.Separable import Mathlib.FieldTheory.NormalClosure -import Mathlib.RingTheory.AlgebraicIndependent +import Mathlib.RingTheory.AlgebraicIndependent.Adjoin +import Mathlib.RingTheory.AlgebraicIndependent.TranscendenceBasis import Mathlib.RingTheory.Polynomial.SeparableDegree import Mathlib.RingTheory.Polynomial.UniqueFactorization diff --git a/Mathlib/RingTheory/AlgebraicIndependent.lean b/Mathlib/RingTheory/AlgebraicIndependent.lean deleted file mode 100644 index 21ed4dbc29d3f..0000000000000 --- a/Mathlib/RingTheory/AlgebraicIndependent.lean +++ /dev/null @@ -1,805 +0,0 @@ -/- -Copyright (c) 2021 Chris Hughes. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Chris Hughes --/ -import Mathlib.Algebra.MvPolynomial.Monad -import Mathlib.Data.Fin.Tuple.Reflection -import Mathlib.FieldTheory.Adjoin -import Mathlib.FieldTheory.MvRatFunc.Rank -import Mathlib.RingTheory.Algebraic.Cardinality -import Mathlib.RingTheory.MvPolynomial.Basic -import Mathlib.RingTheory.Algebraic.MvPolynomial - -/-! -# Algebraic Independence - -This file defines algebraic independence of a family of element of an `R` algebra. - -## Main definitions - -* `AlgebraicIndependent` - `AlgebraicIndependent R x` states the family of elements `x` - is algebraically independent over `R`, meaning that the canonical map out of the multivariable - polynomial ring is injective. - -* `AlgebraicIndependent.aevalEquiv` - The canonical isomorphism from the polynomial ring to the - subalgebra generated by an algebraic independent family. - -* `AlgebraicIndependent.repr` - The canonical map from the subalgebra generated by an - algebraic independent family into the polynomial ring. It is the inverse of - `AlgebraicIndependent.aevalEquiv`. - -* `AlgebraicIndependent.aevalEquivField` - The canonical isomorphism from the rational function - field ring to the intermediate field generated by an algebraic independent family. - -* `AlgebraicIndependent.reprField` - The canonical map from the intermediate field generated by an - algebraic independent family into the rational function field. It is the inverse of - `AlgebraicIndependent.aevalEquivField`. - -* `IsTranscendenceBasis R x` - a family `x` is a transcendence basis over `R` if it is a maximal - algebraically independent subset. - -## References - -* [Stacks: Transcendence](https://stacks.math.columbia.edu/tag/030D) - -## TODO -Define the transcendence degree and show it is independent of the choice of a -transcendence basis. - -## Tags -transcendence basis, transcendence degree, transcendence - --/ - - -noncomputable section - -open Function Set Subalgebra MvPolynomial Algebra - -open scoped Classical - -universe u v w - -variable {ι : Type*} {ι' : Type*} (R : Type*) {K : Type*} -variable {A : Type*} {A' : Type*} -variable (x : ι → A) -variable [CommRing R] [CommRing A] [CommRing A'] [Algebra R A] [Algebra R A'] - -/-- `AlgebraicIndependent R x` states the family of elements `x` - is algebraically independent over `R`, meaning that the canonical - map out of the multivariable polynomial ring is injective. -/ -def AlgebraicIndependent : Prop := - Injective (MvPolynomial.aeval x : MvPolynomial ι R →ₐ[R] A) - -variable {R} {x} - -theorem algebraicIndependent_iff_ker_eq_bot : - AlgebraicIndependent R x ↔ - RingHom.ker (MvPolynomial.aeval x : MvPolynomial ι R →ₐ[R] A).toRingHom = ⊥ := - RingHom.injective_iff_ker_eq_bot _ - -theorem algebraicIndependent_iff : - AlgebraicIndependent R x ↔ - ∀ p : MvPolynomial ι R, MvPolynomial.aeval (x : ι → A) p = 0 → p = 0 := - injective_iff_map_eq_zero _ - -theorem AlgebraicIndependent.eq_zero_of_aeval_eq_zero (h : AlgebraicIndependent R x) : - ∀ p : MvPolynomial ι R, MvPolynomial.aeval (x : ι → A) p = 0 → p = 0 := - algebraicIndependent_iff.1 h - -theorem algebraicIndependent_iff_injective_aeval : - AlgebraicIndependent R x ↔ Injective (MvPolynomial.aeval x : MvPolynomial ι R →ₐ[R] A) := - Iff.rfl - -@[simp] -theorem algebraicIndependent_empty_type_iff [IsEmpty ι] : - AlgebraicIndependent R x ↔ Injective (algebraMap R A) := by - rw [algebraicIndependent_iff_injective_aeval, MvPolynomial.aeval_injective_iff_of_isEmpty] - -/-- A one-element family `x` is algebraically independent if and only if -its element is transcendental. -/ -@[simp] -theorem algebraicIndependent_unique_type_iff [Unique ι] : - AlgebraicIndependent R x ↔ Transcendental R (x default) := by - rw [transcendental_iff_injective, algebraicIndependent_iff_injective_aeval] - let i := (renameEquiv R (Equiv.equivPUnit.{_, 1} ι)).trans (pUnitAlgEquiv R) - have key : aeval (R := R) x = (Polynomial.aeval (R := R) (x default)).comp i := by - ext y - simp [i, Subsingleton.elim y default] - simp [key] - -/-- The one-element family `![x]` is algebraically independent if and only if -`x` is transcendental. -/ -theorem algebraicIndependent_iff_transcendental {x : A} : - AlgebraicIndependent R ![x] ↔ Transcendental R x := by - simp - -namespace AlgebraicIndependent - -theorem of_comp (f : A →ₐ[R] A') (hfv : AlgebraicIndependent R (f ∘ x)) : - AlgebraicIndependent R x := by - have : aeval (f ∘ x) = f.comp (aeval x) := by ext; simp - rw [AlgebraicIndependent, this, AlgHom.coe_comp] at hfv - exact hfv.of_comp - -variable (hx : AlgebraicIndependent R x) -include hx - -theorem algebraMap_injective : Injective (algebraMap R A) := by - simpa [Function.comp_def] using - (Injective.of_comp_iff (algebraicIndependent_iff_injective_aeval.1 hx) MvPolynomial.C).2 - (MvPolynomial.C_injective _ _) - -theorem linearIndependent : LinearIndependent R x := by - rw [linearIndependent_iff_injective_linearCombination] - have : Finsupp.linearCombination R x = - (MvPolynomial.aeval x).toLinearMap.comp (Finsupp.linearCombination R X) := by - ext - simp - rw [this] - refine hx.comp ?_ - rw [← linearIndependent_iff_injective_linearCombination] - exact linearIndependent_X _ _ - -protected theorem injective [Nontrivial R] : Injective x := - hx.linearIndependent.injective - -theorem ne_zero [Nontrivial R] (i : ι) : x i ≠ 0 := - hx.linearIndependent.ne_zero i - -theorem comp (f : ι' → ι) (hf : Function.Injective f) : AlgebraicIndependent R (x ∘ f) := by - intro p q - simpa [aeval_rename, (rename_injective f hf).eq_iff] using @hx (rename f p) (rename f q) - -theorem coe_range : AlgebraicIndependent R ((↑) : range x → A) := by - simpa using hx.comp _ (rangeSplitting_injective x) - -theorem map {f : A →ₐ[R] A'} (hf_inj : Set.InjOn f (adjoin R (range x))) : - AlgebraicIndependent R (f ∘ x) := by - have : aeval (f ∘ x) = f.comp (aeval x) := by ext; simp - have h : ∀ p : MvPolynomial ι R, aeval x p ∈ (@aeval R _ _ _ _ _ ((↑) : range x → A)).range := by - intro p - rw [AlgHom.mem_range] - refine ⟨MvPolynomial.rename (codRestrict x (range x) mem_range_self) p, ?_⟩ - simp [Function.comp_def, aeval_rename] - intro x y hxy - rw [this] at hxy - rw [adjoin_eq_range] at hf_inj - exact hx (hf_inj (h x) (h y) hxy) - -theorem map' {f : A →ₐ[R] A'} (hf_inj : Injective f) : AlgebraicIndependent R (f ∘ x) := - hx.map hf_inj.injOn - -/-- If a family `x` is algebraically independent, then any of its element is transcendental. -/ -theorem transcendental (i : ι) : Transcendental R (x i) := by - have := hx.comp ![i] (Function.injective_of_subsingleton _) - have : AlgebraicIndependent R ![x i] := by rwa [← FinVec.map_eq] at this - rwa [← algebraicIndependent_iff_transcendental] - -/-- If `x = {x_i : A | i : ι}` and `f = {f_i : MvPolynomial ι R | i : ι}` are algebraically -independent over `R`, then `{f_i(x) | i : ι}` is also algebraically independent over `R`. -For the partial converse, see `AlgebraicIndependent.of_aeval`. -/ -theorem aeval_of_algebraicIndependent - {f : ι → MvPolynomial ι R} (hf : AlgebraicIndependent R f) : - AlgebraicIndependent R fun i ↦ aeval x (f i) := by - rw [algebraicIndependent_iff] at hx hf ⊢ - intro p hp - exact hf _ (hx _ (by rwa [← aeval_comp_bind₁, AlgHom.comp_apply] at hp)) - -omit hx in -/-- If `{f_i(x) | i : ι}` is algebraically independent over `R`, then -`{f_i : MvPolynomial ι R | i : ι}` is also algebraically independent over `R`. -In fact, the `x = {x_i : A | i : ι}` is also transcendental over `R` provided that `R` -is a field and `ι` is finite; the proof needs transcendence degree. -/ -theorem of_aeval {f : ι → MvPolynomial ι R} - (H : AlgebraicIndependent R fun i ↦ aeval x (f i)) : - AlgebraicIndependent R f := by - rw [algebraicIndependent_iff] at H ⊢ - intro p hp - exact H p (by rw [← aeval_comp_bind₁, AlgHom.comp_apply, bind₁, hp, map_zero]) - -/-- If `A/R` is algebraic, then all algebraically independent families are empty. -/ -theorem isEmpty_of_isAlgebraic [Algebra.IsAlgebraic R A] : IsEmpty ι := by - rcases isEmpty_or_nonempty ι with h | ⟨⟨i⟩⟩ - · exact h - exact False.elim (hx.transcendental i (Algebra.IsAlgebraic.isAlgebraic _)) - -end AlgebraicIndependent - -theorem MvPolynomial.algebraicIndependent_X (σ R : Type*) [CommRing R] : - AlgebraicIndependent R (X (R := R) (σ := σ)) := by - rw [AlgebraicIndependent, aeval_X_left] - exact injective_id - -open AlgebraicIndependent - -theorem AlgHom.algebraicIndependent_iff (f : A →ₐ[R] A') (hf : Injective f) : - AlgebraicIndependent R (f ∘ x) ↔ AlgebraicIndependent R x := - ⟨fun h => h.of_comp f, fun h => h.map hf.injOn⟩ - -@[nontriviality] -theorem algebraicIndependent_of_subsingleton [Subsingleton R] : AlgebraicIndependent R x := - algebraicIndependent_iff.2 fun _ _ => Subsingleton.elim _ _ - -theorem algebraicIndependent_equiv (e : ι ≃ ι') {f : ι' → A} : - AlgebraicIndependent R (f ∘ e) ↔ AlgebraicIndependent R f := - ⟨fun h => Function.comp_id f ▸ e.self_comp_symm ▸ h.comp _ e.symm.injective, - fun h => h.comp _ e.injective⟩ - -theorem algebraicIndependent_equiv' (e : ι ≃ ι') {f : ι' → A} {g : ι → A} (h : f ∘ e = g) : - AlgebraicIndependent R g ↔ AlgebraicIndependent R f := - h ▸ algebraicIndependent_equiv e - -theorem algebraicIndependent_subtype_range {ι} {f : ι → A} (hf : Injective f) : - AlgebraicIndependent R ((↑) : range f → A) ↔ AlgebraicIndependent R f := - Iff.symm <| algebraicIndependent_equiv' (Equiv.ofInjective f hf) rfl - -alias ⟨AlgebraicIndependent.of_subtype_range, _⟩ := algebraicIndependent_subtype_range - -theorem algebraicIndependent_image {ι} {s : Set ι} {f : ι → A} (hf : Set.InjOn f s) : - (AlgebraicIndependent R fun x : s => f x) ↔ AlgebraicIndependent R fun x : f '' s => (x : A) := - algebraicIndependent_equiv' (Equiv.Set.imageOfInjOn _ _ hf) rfl - -theorem algebraicIndependent_adjoin (hs : AlgebraicIndependent R x) : - @AlgebraicIndependent ι R (adjoin R (range x)) - (fun i : ι => ⟨x i, subset_adjoin (mem_range_self i)⟩) _ _ _ := - AlgebraicIndependent.of_comp (adjoin R (range x)).val hs - -/-- A set of algebraically independent elements in an algebra `A` over a ring `K` is also -algebraically independent over a subring `R` of `K`. -/ -theorem AlgebraicIndependent.restrictScalars {K : Type*} [CommRing K] [Algebra R K] [Algebra K A] - [IsScalarTower R K A] (hinj : Function.Injective (algebraMap R K)) - (ai : AlgebraicIndependent K x) : AlgebraicIndependent R x := by - have : (aeval x : MvPolynomial ι K →ₐ[K] A).toRingHom.comp (MvPolynomial.map (algebraMap R K)) = - (aeval x : MvPolynomial ι R →ₐ[R] A).toRingHom := by - ext <;> simp [algebraMap_eq_smul_one] - show Injective (aeval x).toRingHom - rw [← this, RingHom.coe_comp] - exact Injective.comp ai (MvPolynomial.map_injective _ hinj) - -section RingHom - -variable {S B FRS FAB : Type*} [CommRing S] [CommRing B] [Algebra S B] - -section - -variable [FunLike FRS R S] [RingHomClass FRS R S] [FunLike FAB A B] [RingHomClass FAB A B] - (f : FRS) (g : FAB) - -theorem AlgebraicIndependent.of_ringHom_of_comp_eq (H : AlgebraicIndependent S (g ∘ x)) - (hf : Function.Injective f) - (h : RingHom.comp (algebraMap S B) f = RingHom.comp g (algebraMap R A)) : - AlgebraicIndependent R x := by - rw [algebraicIndependent_iff] at H ⊢ - intro p hp - have := H (p.map f) <| by - have : (g : A →+* B) _ = _ := congr(g $hp) - rwa [map_zero, map_aeval, ← h, ← eval₂Hom_map_hom, ← aeval_eq_eval₂Hom] at this - exact map_injective (f : R →+* S) hf (by rwa [map_zero]) - -theorem AlgebraicIndependent.ringHom_of_comp_eq (H : AlgebraicIndependent R x) - (hf : Function.Surjective f) (hg : Function.Injective g) - (h : RingHom.comp (algebraMap S B) f = RingHom.comp g (algebraMap R A)) : - AlgebraicIndependent S (g ∘ x) := by - rw [algebraicIndependent_iff] at H ⊢ - intro p hp - obtain ⟨q, rfl⟩ := map_surjective (f : R →+* S) hf p - rw [H q (hg (by rwa [map_zero, ← RingHom.coe_coe g, map_aeval, ← h, ← eval₂Hom_map_hom, - ← aeval_eq_eval₂Hom])), map_zero] - -end - -section - -variable [EquivLike FRS R S] [RingEquivClass FRS R S] [FunLike FAB A B] [RingHomClass FAB A B] - (f : FRS) (g : FAB) - -theorem algebraicIndependent_ringHom_iff_of_comp_eq - (hg : Function.Injective g) - (h : RingHom.comp (algebraMap S B) f = RingHom.comp g (algebraMap R A)) : - AlgebraicIndependent S (g ∘ x) ↔ AlgebraicIndependent R x := - ⟨fun H ↦ H.of_ringHom_of_comp_eq f g (EquivLike.injective f) h, - fun H ↦ H.ringHom_of_comp_eq f g (EquivLike.surjective f) hg h⟩ - -end - -end RingHom - -/-- Every finite subset of an algebraically independent set is algebraically independent. -/ -theorem algebraicIndependent_finset_map_embedding_subtype (s : Set A) - (li : AlgebraicIndependent R ((↑) : s → A)) (t : Finset s) : - AlgebraicIndependent R ((↑) : Finset.map (Embedding.subtype s) t → A) := by - let f : t.map (Embedding.subtype s) → s := fun x => - ⟨x.1, by - obtain ⟨x, h⟩ := x - rw [Finset.mem_map] at h - obtain ⟨a, _, rfl⟩ := h - simp only [Subtype.coe_prop, Embedding.coe_subtype]⟩ - convert AlgebraicIndependent.comp li f _ - rintro ⟨x, hx⟩ ⟨y, hy⟩ - rw [Finset.mem_map] at hx hy - obtain ⟨a, _, rfl⟩ := hx - obtain ⟨b, _, rfl⟩ := hy - simp only [f, imp_self, Subtype.mk_eq_mk] - -/-- If every finite set of algebraically independent element has cardinality at most `n`, -then the same is true for arbitrary sets of algebraically independent elements. -/ -theorem algebraicIndependent_bounded_of_finset_algebraicIndependent_bounded {n : ℕ} - (H : ∀ s : Finset A, (AlgebraicIndependent R fun i : s => (i : A)) → s.card ≤ n) : - ∀ s : Set A, AlgebraicIndependent R ((↑) : s → A) → Cardinal.mk s ≤ n := by - intro s li - apply Cardinal.card_le_of - intro t - rw [← Finset.card_map (Embedding.subtype s)] - apply H - apply algebraicIndependent_finset_map_embedding_subtype _ li - -section Subtype - -theorem AlgebraicIndependent.restrict_of_comp_subtype {s : Set ι} - (hs : AlgebraicIndependent R (x ∘ (↑) : s → A)) : AlgebraicIndependent R (s.restrict x) := - hs - -variable (R A) - -theorem algebraicIndependent_empty_iff : - AlgebraicIndependent R ((↑) : (∅ : Set A) → A) ↔ Injective (algebraMap R A) := by simp - -variable {R A} - -theorem AlgebraicIndependent.mono {t s : Set A} (h : t ⊆ s) - (hx : AlgebraicIndependent R ((↑) : s → A)) : AlgebraicIndependent R ((↑) : t → A) := by - simpa [Function.comp] using hx.comp (inclusion h) (inclusion_injective h) - -end Subtype - -theorem AlgebraicIndependent.to_subtype_range {ι} {f : ι → A} (hf : AlgebraicIndependent R f) : - AlgebraicIndependent R ((↑) : range f → A) := by - nontriviality R - rwa [algebraicIndependent_subtype_range hf.injective] - -theorem AlgebraicIndependent.to_subtype_range' {ι} {f : ι → A} (hf : AlgebraicIndependent R f) {t} - (ht : range f = t) : AlgebraicIndependent R ((↑) : t → A) := - ht ▸ hf.to_subtype_range - -theorem algebraicIndependent_comp_subtype {s : Set ι} : - AlgebraicIndependent R (x ∘ (↑) : s → A) ↔ - ∀ p ∈ MvPolynomial.supported R s, aeval x p = 0 → p = 0 := by - have : (aeval (x ∘ (↑) : s → A) : _ →ₐ[R] _) = (aeval x).comp (rename (↑)) := by ext; simp - have : ∀ p : MvPolynomial s R, rename ((↑) : s → ι) p = 0 ↔ p = 0 := - (injective_iff_map_eq_zero' (rename ((↑) : s → ι) : MvPolynomial s R →ₐ[R] _).toRingHom).1 - (rename_injective _ Subtype.val_injective) - simp [algebraicIndependent_iff, supported_eq_range_rename, *] - -theorem algebraicIndependent_subtype {s : Set A} : - AlgebraicIndependent R ((↑) : s → A) ↔ - ∀ p : MvPolynomial A R, p ∈ MvPolynomial.supported R s → aeval id p = 0 → p = 0 := by - apply @algebraicIndependent_comp_subtype _ _ _ id - -theorem algebraicIndependent_of_finite (s : Set A) - (H : ∀ t ⊆ s, t.Finite → AlgebraicIndependent R ((↑) : t → A)) : - AlgebraicIndependent R ((↑) : s → A) := - algebraicIndependent_subtype.2 fun p hp => - algebraicIndependent_subtype.1 (H _ (mem_supported.1 hp) (Finset.finite_toSet _)) _ (by simp) - -theorem AlgebraicIndependent.image_of_comp {ι ι'} (s : Set ι) (f : ι → ι') (g : ι' → A) - (hs : AlgebraicIndependent R fun x : s => g (f x)) : - AlgebraicIndependent R fun x : f '' s => g x := by - nontriviality R - have : InjOn f s := injOn_iff_injective.2 hs.injective.of_comp - exact (algebraicIndependent_equiv' (Equiv.Set.imageOfInjOn f s this) rfl).1 hs - -theorem AlgebraicIndependent.image {ι} {s : Set ι} {f : ι → A} - (hs : AlgebraicIndependent R fun x : s => f x) : - AlgebraicIndependent R fun x : f '' s => (x : A) := by - convert AlgebraicIndependent.image_of_comp s f id hs - -theorem algebraicIndependent_iUnion_of_directed {η : Type*} [Nonempty η] {s : η → Set A} - (hs : Directed (· ⊆ ·) s) (h : ∀ i, AlgebraicIndependent R ((↑) : s i → A)) : - AlgebraicIndependent R ((↑) : (⋃ i, s i) → A) := by - refine algebraicIndependent_of_finite (⋃ i, s i) fun t ht ft => ?_ - rcases finite_subset_iUnion ft ht with ⟨I, fi, hI⟩ - rcases hs.finset_le fi.toFinset with ⟨i, hi⟩ - exact (h i).mono (Subset.trans hI <| iUnion₂_subset fun j hj => hi j (fi.mem_toFinset.2 hj)) - -theorem algebraicIndependent_sUnion_of_directed {s : Set (Set A)} (hsn : s.Nonempty) - (hs : DirectedOn (· ⊆ ·) s) (h : ∀ a ∈ s, AlgebraicIndependent R ((↑) : a → A)) : - AlgebraicIndependent R ((↑) : ⋃₀ s → A) := by - letI : Nonempty s := Nonempty.to_subtype hsn - rw [sUnion_eq_iUnion] - exact algebraicIndependent_iUnion_of_directed hs.directed_val (by simpa using h) - -theorem exists_maximal_algebraicIndependent (s t : Set A) (hst : s ⊆ t) - (hs : AlgebraicIndependent R ((↑) : s → A)) : ∃ u, s ⊆ u ∧ - Maximal (fun (x : Set A) ↦ AlgebraicIndependent R ((↑) : x → A) ∧ x ⊆ t) u := by - refine zorn_subset_nonempty { u : Set A | AlgebraicIndependent R ((↑) : u → A) ∧ u ⊆ t} - (fun c hc chainc hcn ↦ ⟨⋃₀ c, ⟨?_, ?_⟩, fun _ ↦ subset_sUnion_of_mem⟩) s ⟨hs, hst⟩ - · exact algebraicIndependent_sUnion_of_directed hcn chainc.directedOn (fun x hxc ↦ (hc hxc).1) - exact fun x ⟨w, hyc, hwy⟩ ↦ (hc hyc).2 hwy - -namespace AlgebraicIndependent - -section repr - -variable (hx : AlgebraicIndependent R x) -include hx - -/-- Canonical isomorphism between polynomials and the subalgebra generated by - algebraically independent elements. -/ -@[simps! apply_coe] -def aevalEquiv : MvPolynomial ι R ≃ₐ[R] Algebra.adjoin R (range x) := - (AlgEquiv.ofInjective (aeval x) (algebraicIndependent_iff_injective_aeval.1 hx)).trans - (Subalgebra.equivOfEq _ _ (Algebra.adjoin_range_eq_range_aeval R x).symm) - ---@[simp] Porting note: removing simp because the linter complains about deterministic timeout -theorem algebraMap_aevalEquiv (p : MvPolynomial ι R) : - algebraMap (Algebra.adjoin R (range x)) A (hx.aevalEquiv p) = aeval x p := - rfl - -/-- The canonical map from the subalgebra generated by an algebraic independent family - into the polynomial ring. -/ -def repr : Algebra.adjoin R (range x) →ₐ[R] MvPolynomial ι R := - hx.aevalEquiv.symm - -@[simp] -theorem aeval_repr (p) : aeval x (hx.repr p) = p := - Subtype.ext_iff.1 (AlgEquiv.apply_symm_apply hx.aevalEquiv p) - -theorem aeval_comp_repr : (aeval x).comp hx.repr = Subalgebra.val _ := - AlgHom.ext <| hx.aeval_repr - -theorem repr_ker : RingHom.ker (hx.repr : adjoin R (range x) →+* MvPolynomial ι R) = ⊥ := - (RingHom.injective_iff_ker_eq_bot _).1 (AlgEquiv.injective _) - -end repr - -section reprField - -variable {F E : Type*} {x : ι → E} [Field F] [Field E] [Algebra F E] (hx : AlgebraicIndependent F x) -include hx - -/-- Canonical isomorphism between rational function field and the - intermediate field generated by algebraically independent elements. -/ -def aevalEquivField : - FractionRing (MvPolynomial ι F) ≃ₐ[F] ↥(IntermediateField.adjoin F (range x)) := - let i := IsFractionRing.liftAlgHom (K := FractionRing (MvPolynomial ι F)) - (algebraicIndependent_iff_injective_aeval.2 hx) - (show _ ≃ₐ[F] i.fieldRange from AlgEquiv.ofInjectiveField i).trans <| - IntermediateField.equivOfEq <| - IsFractionRing.algHom_fieldRange_eq_of_comp_eq_of_range_eq (g := aeval x) (f := i) - (by ext <;> simp [i]) (Algebra.adjoin_range_eq_range_aeval F x).symm - -@[simp] -theorem aevalEquivField_apply_coe (a : FractionRing (MvPolynomial ι F)) : - hx.aevalEquivField a = - IsFractionRing.lift (algebraicIndependent_iff_injective_aeval.2 hx) a := rfl - -theorem aevalEquivField_algebraMap_apply_coe (a : MvPolynomial ι F) : - hx.aevalEquivField (algebraMap _ _ a) = aeval x a := by - simp - -/-- The canonical map from the intermediate field generated by an algebraic independent family - into the rational function field. -/ -def reprField : IntermediateField.adjoin F (range x) →ₐ[F] FractionRing (MvPolynomial ι F) := - hx.aevalEquivField.symm - -@[simp] -theorem lift_reprField (p) : - IsFractionRing.lift (algebraicIndependent_iff_injective_aeval.2 hx) (hx.reprField p) = p := - Subtype.ext_iff.1 (AlgEquiv.apply_symm_apply hx.aevalEquivField p) - -theorem liftAlgHom_comp_reprField : - (IsFractionRing.liftAlgHom (algebraicIndependent_iff_injective_aeval.2 hx)).comp hx.reprField = - IntermediateField.val _ := - AlgHom.ext <| hx.lift_reprField - -end reprField - -end AlgebraicIndependent - --- TODO - make this an `AlgEquiv` -/-- The isomorphism between `MvPolynomial (Option ι) R` and the polynomial ring over -the algebra generated by an algebraically independent family. -/ -def AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin (hx : AlgebraicIndependent R x) : - MvPolynomial (Option ι) R ≃+* Polynomial (adjoin R (Set.range x)) := - (MvPolynomial.optionEquivLeft _ _).toRingEquiv.trans - (Polynomial.mapEquiv hx.aevalEquiv.toRingEquiv) - -@[simp] -theorem AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_apply - (hx : AlgebraicIndependent R x) (y) : - hx.mvPolynomialOptionEquivPolynomialAdjoin y = - Polynomial.map (hx.aevalEquiv : MvPolynomial ι R →+* adjoin R (range x)) - (aeval (fun o : Option ι => o.elim Polynomial.X fun s : ι => Polynomial.C (X s)) y) := - rfl - -/-- `simp`-normal form of `mvPolynomialOptionEquivPolynomialAdjoin_C` -/ -@[simp] -theorem AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_C' - (hx : AlgebraicIndependent R x) (r) : - Polynomial.C (hx.aevalEquiv (C r)) = Polynomial.C (algebraMap _ _ r) := by - congr - apply_fun Subtype.val using Subtype.val_injective - simp - -theorem AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_C - (hx : AlgebraicIndependent R x) (r) : - hx.mvPolynomialOptionEquivPolynomialAdjoin (C r) = Polynomial.C (algebraMap _ _ r) := by - simp - -theorem AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_X_none - (hx : AlgebraicIndependent R x) : - hx.mvPolynomialOptionEquivPolynomialAdjoin (X none) = Polynomial.X := by - rw [AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_apply, aeval_X, Option.elim, - Polynomial.map_X] - -theorem AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_X_some - (hx : AlgebraicIndependent R x) (i) : - hx.mvPolynomialOptionEquivPolynomialAdjoin (X (some i)) = - Polynomial.C (hx.aevalEquiv (X i)) := by - rw [AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_apply, aeval_X, Option.elim, - Polynomial.map_C, RingHom.coe_coe] - -theorem AlgebraicIndependent.aeval_comp_mvPolynomialOptionEquivPolynomialAdjoin - (hx : AlgebraicIndependent R x) (a : A) : - RingHom.comp - (↑(Polynomial.aeval a : Polynomial (adjoin R (Set.range x)) →ₐ[_] A) : - Polynomial (adjoin R (Set.range x)) →+* A) - hx.mvPolynomialOptionEquivPolynomialAdjoin.toRingHom = - ↑(MvPolynomial.aeval fun o : Option ι => o.elim a x : MvPolynomial (Option ι) R →ₐ[R] A) := by - refine MvPolynomial.ringHom_ext ?_ ?_ <;> - simp only [RingHom.comp_apply, RingEquiv.toRingHom_eq_coe, RingEquiv.coe_toRingHom, - AlgHom.coe_toRingHom, AlgHom.coe_toRingHom] - · intro r - rw [hx.mvPolynomialOptionEquivPolynomialAdjoin_C, aeval_C, Polynomial.aeval_C, - IsScalarTower.algebraMap_apply R (adjoin R (range x)) A] - · rintro (⟨⟩ | ⟨i⟩) - · rw [hx.mvPolynomialOptionEquivPolynomialAdjoin_X_none, aeval_X, Polynomial.aeval_X, - Option.elim] - · rw [hx.mvPolynomialOptionEquivPolynomialAdjoin_X_some, Polynomial.aeval_C, - hx.algebraMap_aevalEquiv, aeval_X, aeval_X, Option.elim] - -theorem AlgebraicIndependent.option_iff (hx : AlgebraicIndependent R x) (a : A) : - (AlgebraicIndependent R fun o : Option ι => o.elim a x) ↔ - Transcendental (adjoin R (Set.range x)) a := by - rw [algebraicIndependent_iff_injective_aeval, transcendental_iff_injective, - ← AlgHom.coe_toRingHom, ← hx.aeval_comp_mvPolynomialOptionEquivPolynomialAdjoin, - RingHom.coe_comp] - exact Injective.of_comp_iff' (Polynomial.aeval a) - (mvPolynomialOptionEquivPolynomialAdjoin hx).bijective - -/-- Variant of `algebraicIndependent_of_finite` using `Transcendental`. -/ -theorem algebraicIndependent_of_finite' (s : Set A) - (hinj : Injective (algebraMap R A)) - (H : ∀ t ⊆ s, t.Finite → ∀ a ∈ s, a ∉ t → Transcendental (adjoin R t) a) : - AlgebraicIndependent R ((↑) : s → A) := by - classical - refine algebraicIndependent_of_finite s fun t hts hfin ↦ hfin.induction_on' - ((algebraicIndependent_empty_iff R A).2 hinj) fun {a} {u} ha hu ha' h ↦ ?_ - convert ((Subtype.range_coe ▸ h.option_iff a).2 <| H u (hu.trans hts) (hfin.subset hu) - a (hts ha) ha').comp _ (Set.subtypeInsertEquivOption ha').injective - ext x - by_cases h : ↑x = a <;> simp [h, Set.subtypeInsertEquivOption] - -/-- `Type` version of `algebraicIndependent_of_finite'`. -/ -theorem algebraicIndependent_of_finite_type' - (hinj : Injective (algebraMap R A)) - (H : ∀ t : Set ι, t.Finite → ∀ i : ι, i ∉ t → Transcendental (adjoin R (x '' t)) (x i)) : - AlgebraicIndependent R x := by - nontriviality R - haveI := hinj.nontrivial - have hx : Injective x := by - simp_rw [Transcendental] at H - contrapose! H - obtain ⟨i, j, h1, h2⟩ := not_injective_iff.1 H - refine ⟨{j}, by simp, i, by simp [h2], ?_⟩ - rw [h1, Set.image_singleton] - exact isAlgebraic_algebraMap (⟨x j, Algebra.self_mem_adjoin_singleton R _⟩ : adjoin R {x j}) - rw [← Set.coe_comp_rangeFactorization x] - refine .comp (algebraicIndependent_of_finite' _ hinj fun t ht hfin a ha ha' ↦ ?_) _ - (Set.rightInverse_rangeSplitting hx).injective - change Finite t at hfin - have := H (x ⁻¹' t) (Finite.of_injective _ (hx.restrictPreimage t)) - ((Equiv.ofInjective _ hx).symm ⟨_, ha⟩) - (by rwa [Set.mem_preimage, Equiv.apply_ofInjective_symm hx]) - rwa [Set.image_preimage_eq_inter_range, Set.inter_eq_self_of_subset_left ht, - Equiv.apply_ofInjective_symm hx] at this - -namespace MvPolynomial - -/-- If for each `i : ι`, `f_i : R[X]` is transcendental over `R`, then `{f_i(X_i) | i : ι}` -in `MvPolynomial ι R` is algebraically independent over `R`. -/ -theorem algebraicIndependent_polynomial_aeval_X - (f : ι → Polynomial R) (hf : ∀ i, Transcendental R (f i)) : - AlgebraicIndependent R fun i ↦ Polynomial.aeval (X i : MvPolynomial ι R) (f i) := by - set x := fun i ↦ Polynomial.aeval (X i : MvPolynomial ι R) (f i) - refine algebraicIndependent_of_finite_type' (C_injective _ _) fun t _ i hi ↦ ?_ - have hle : adjoin R (x '' t) ≤ supported R t := by - rw [Algebra.adjoin_le_iff, Set.image_subset_iff] - intro _ h - rw [Set.mem_preimage] - refine Algebra.adjoin_mono ?_ (Polynomial.aeval_mem_adjoin_singleton R _) - simp_rw [singleton_subset_iff, Set.mem_image_of_mem _ h] - exact (transcendental_supported_polynomial_aeval_X R hi (hf i)).of_tower_top_of_subalgebra_le hle - -end MvPolynomial - -/-- If `{x_i : A | i : ι}` is algebraically independent over `R`, and for each `i`, -`f_i : R[X]` is transcendental over `R`, then `{f_i(x_i) | i : ι}` is also -algebraically independent over `R`. -/ -theorem AlgebraicIndependent.polynomial_aeval_of_transcendental - (hx : AlgebraicIndependent R x) - {f : ι → Polynomial R} (hf : ∀ i, Transcendental R (f i)) : - AlgebraicIndependent R fun i ↦ Polynomial.aeval (x i) (f i) := by - convert aeval_of_algebraicIndependent hx (algebraicIndependent_polynomial_aeval_X _ hf) - rw [← AlgHom.comp_apply] - congr 1; ext1; simp - -variable (R) - -/-- A family is a transcendence basis if it is a maximal algebraically independent subset. -/ -def IsTranscendenceBasis (x : ι → A) : Prop := - AlgebraicIndependent R x ∧ - ∀ (s : Set A) (_ : AlgebraicIndependent R ((↑) : s → A)) (_ : range x ≤ s), range x = s - -theorem exists_isTranscendenceBasis (h : Injective (algebraMap R A)) : - ∃ s : Set A, IsTranscendenceBasis R ((↑) : s → A) := by - cases' exists_maximal_algebraicIndependent (∅ : Set A) Set.univ (Set.subset_univ _) - ((algebraicIndependent_empty_iff R A).2 h) with - s hs - refine ⟨s, hs.2.1.1, fun t ht hst ↦ ?_⟩ - simp only [Subtype.range_coe_subtype, setOf_mem_eq] at * - exact hs.2.eq_of_le ⟨ht, subset_univ _⟩ hst - -/-- `Type` version of `exists_isTranscendenceBasis`. -/ -theorem exists_isTranscendenceBasis' (R : Type u) {A : Type v} [CommRing R] [CommRing A] - [Algebra R A] (h : Injective (algebraMap R A)) : - ∃ (ι : Type v) (x : ι → A), IsTranscendenceBasis R x := by - obtain ⟨s, h⟩ := exists_isTranscendenceBasis R h - exact ⟨s, Subtype.val, h⟩ - -variable {R} - -theorem AlgebraicIndependent.isTranscendenceBasis_iff {ι : Type w} {R : Type u} [CommRing R] - [Nontrivial R] {A : Type v} [CommRing A] [Algebra R A] {x : ι → A} - (i : AlgebraicIndependent R x) : - IsTranscendenceBasis R x ↔ - ∀ (κ : Type v) (w : κ → A) (_ : AlgebraicIndependent R w) (j : ι → κ) (_ : w ∘ j = x), - Surjective j := by - fconstructor - · rintro p κ w i' j rfl - have p := p.2 (range w) i'.coe_range (range_comp_subset_range _ _) - rw [range_comp, ← @image_univ _ _ w] at p - exact range_eq_univ.mp (image_injective.mpr i'.injective p) - · intro p - use i - intro w i' h - specialize p w ((↑) : w → A) i' (fun i => ⟨x i, range_subset_iff.mp h i⟩) (by ext; simp) - have q := congr_arg (fun s => ((↑) : w → A) '' s) p.range_eq - dsimp at q - rw [← image_univ, image_image] at q - simpa using q - -theorem IsTranscendenceBasis.isAlgebraic [Nontrivial R] (hx : IsTranscendenceBasis R x) : - Algebra.IsAlgebraic (adjoin R (range x)) A := by - constructor - intro a - rw [← not_iff_comm.1 (hx.1.option_iff _).symm] - intro ai - have h₁ : range x ⊆ range fun o : Option ι => o.elim a x := by - rintro x ⟨y, rfl⟩ - exact ⟨some y, rfl⟩ - have h₂ : range x ≠ range fun o : Option ι => o.elim a x := by - intro h - have : a ∈ range x := by - rw [h] - exact ⟨none, rfl⟩ - rcases this with ⟨b, rfl⟩ - have : some b = none := ai.injective rfl - simpa - exact h₂ (hx.2 (Set.range fun o : Option ι => o.elim a x) - ((algebraicIndependent_subtype_range ai.injective).2 ai) h₁) - -/-- If `x` is a transcendence basis of `A/R`, then it is empty if and only if -`A/R` is algebraic. -/ -theorem IsTranscendenceBasis.isEmpty_iff_isAlgebraic [Nontrivial R] - (hx : IsTranscendenceBasis R x) : - IsEmpty ι ↔ Algebra.IsAlgebraic R A := by - refine ⟨fun _ ↦ ?_, fun _ ↦ hx.1.isEmpty_of_isAlgebraic⟩ - have := hx.isAlgebraic - rw [Set.range_eq_empty x, adjoin_empty] at this - exact algebra_isAlgebraic_of_algebra_isAlgebraic_bot_left R A - -/-- If `x` is a transcendence basis of `A/R`, then it is not empty if and only if -`A/R` is transcendental. -/ -theorem IsTranscendenceBasis.nonempty_iff_transcendental [Nontrivial R] - (hx : IsTranscendenceBasis R x) : - Nonempty ι ↔ Algebra.Transcendental R A := by - rw [← not_isEmpty_iff, Algebra.transcendental_iff_not_isAlgebraic, hx.isEmpty_iff_isAlgebraic] - -theorem IsTranscendenceBasis.isAlgebraic_field {F E : Type*} {x : ι → E} - [Field F] [Field E] [Algebra F E] (hx : IsTranscendenceBasis F x) : - Algebra.IsAlgebraic (IntermediateField.adjoin F (range x)) E := by - haveI := hx.isAlgebraic - set S := range x - letI : Algebra (adjoin F S) (IntermediateField.adjoin F S) := - (Subalgebra.inclusion (IntermediateField.algebra_adjoin_le_adjoin F S)).toRingHom.toAlgebra - haveI : IsScalarTower (adjoin F S) (IntermediateField.adjoin F S) E := - IsScalarTower.of_algebraMap_eq (congrFun rfl) - exact Algebra.IsAlgebraic.extendScalars (R := adjoin F S) (Subalgebra.inclusion_injective _) - -section Field - -variable [Field K] [Algebra K A] - -/- Porting note: removing `simp`, not in simp normal form. Could make `Function.Injective f` a -simp lemma when `f` is a field hom, and then simp would prove this -/ -theorem algebraicIndependent_empty_type [IsEmpty ι] [Nontrivial A] : AlgebraicIndependent K x := by - rw [algebraicIndependent_empty_type_iff] - exact RingHom.injective _ - -theorem algebraicIndependent_empty [Nontrivial A] : - AlgebraicIndependent K ((↑) : (∅ : Set A) → A) := - algebraicIndependent_empty_type - -end Field - -section RankAndCardinality - -open Cardinal - -theorem IsTranscendenceBasis.lift_cardinalMk_eq_max_lift - {F : Type u} {E : Type v} [CommRing F] [Nontrivial F] [CommRing E] [IsDomain E] [Algebra F E] - {ι : Type w} {x : ι → E} [Nonempty ι] (hx : IsTranscendenceBasis F x) : - lift.{max u w} #E = lift.{max v w} #F ⊔ lift.{max u v} #ι ⊔ ℵ₀ := by - let K := Algebra.adjoin F (Set.range x) - suffices #E = #K by simp [this, ← lift_mk_eq'.2 ⟨hx.1.aevalEquiv.toEquiv⟩] - haveI : Algebra.IsAlgebraic K E := hx.isAlgebraic - refine le_antisymm ?_ (mk_le_of_injective Subtype.val_injective) - haveI : Infinite K := hx.1.aevalEquiv.infinite_iff.1 inferInstance - simpa only [sup_eq_left.2 (aleph0_le_mk K)] using Algebra.IsAlgebraic.cardinalMk_le_max K E - -theorem IsTranscendenceBasis.lift_rank_eq_max_lift - {F : Type u} {E : Type v} [Field F] [Field E] [Algebra F E] - {ι : Type w} {x : ι → E} [Nonempty ι] (hx : IsTranscendenceBasis F x) : - lift.{max u w} (Module.rank F E) = lift.{max v w} #F ⊔ lift.{max u v} #ι ⊔ ℵ₀ := by - let K := IntermediateField.adjoin F (Set.range x) - haveI : Algebra.IsAlgebraic K E := hx.isAlgebraic_field - rw [← rank_mul_rank F K E, lift_mul, ← hx.1.aevalEquivField.toLinearEquiv.lift_rank_eq, - MvRatFunc.rank_eq_max_lift, lift_max, lift_max, lift_lift, lift_lift, lift_aleph0] - refine mul_eq_left le_sup_right ((lift_le.2 ((rank_le_card K E).trans - (Algebra.IsAlgebraic.cardinalMk_le_max K E))).trans_eq ?_) (by simp [rank_pos.ne']) - simp [← lift_mk_eq'.2 ⟨hx.1.aevalEquivField.toEquiv⟩] - -theorem Algebra.Transcendental.rank_eq_cardinalMk - (F : Type u) (E : Type v) [Field F] [Field E] [Algebra F E] [Algebra.Transcendental F E] : - Module.rank F E = #E := by - obtain ⟨ι, x, hx⟩ := exists_isTranscendenceBasis' _ (algebraMap F E).injective - haveI := hx.nonempty_iff_transcendental.2 ‹_› - simpa [← hx.lift_cardinalMk_eq_max_lift] using hx.lift_rank_eq_max_lift - -theorem IntermediateField.rank_sup_le - {F : Type u} {E : Type v} [Field F] [Field E] [Algebra F E] (A B : IntermediateField F E) : - Module.rank F ↥(A ⊔ B) ≤ Module.rank F A * Module.rank F B := by - by_cases hA : Algebra.IsAlgebraic F A - · exact rank_sup_le_of_isAlgebraic A B (Or.inl hA) - by_cases hB : Algebra.IsAlgebraic F B - · exact rank_sup_le_of_isAlgebraic A B (Or.inr hB) - rw [← Algebra.transcendental_iff_not_isAlgebraic] at hA hB - haveI : Algebra.Transcendental F ↥(A ⊔ B) := .ringHom_of_comp_eq (RingHom.id F) - (inclusion le_sup_left) Function.surjective_id (inclusion_injective _) rfl - haveI := Algebra.Transcendental.infinite F A - haveI := Algebra.Transcendental.infinite F B - simp_rw [Algebra.Transcendental.rank_eq_cardinalMk] - rw [sup_def, mul_mk_eq_max, ← Cardinal.lift_le.{u}] - refine (lift_cardinalMk_adjoin_le _ _).trans ?_ - calc - _ ≤ Cardinal.lift.{v} #F ⊔ Cardinal.lift.{u} (#A ⊔ #B) ⊔ ℵ₀ := by - gcongr - rw [Cardinal.lift_le] - exact (mk_union_le _ _).trans_eq (by simp) - _ = _ := by - simp [lift_mk_le_lift_mk_of_injective (algebraMap F A).injective] - -end RankAndCardinality diff --git a/Mathlib/RingTheory/AlgebraicIndependent/Adjoin.lean b/Mathlib/RingTheory/AlgebraicIndependent/Adjoin.lean new file mode 100644 index 0000000000000..53d2a48d0cf9b --- /dev/null +++ b/Mathlib/RingTheory/AlgebraicIndependent/Adjoin.lean @@ -0,0 +1,71 @@ +/- +Copyright (c) 2021 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes +-/ +import Mathlib.FieldTheory.Adjoin +import Mathlib.RingTheory.AlgebraicIndependent.Defs + +/-! +# Algebraic Independence + +This file concerns adjoining an algebraic independent family to a field. + +## Main definitions + +* `AlgebraicIndependent.aevalEquivField` - The canonical isomorphism from the rational function + field ring to the intermediate field generated by an algebraic independent family. + +* `AlgebraicIndependent.reprField` - The canonical map from the intermediate field generated by an + algebraic independent family into the rational function field. It is the inverse of + `AlgebraicIndependent.aevalEquivField`. +-/ + +noncomputable section + +open Function Set Subalgebra MvPolynomial Algebra + +open scoped Classical + +namespace AlgebraicIndependent + +variable {ι : Type*} +variable {F E : Type*} {x : ι → E} [Field F] [Field E] [Algebra F E] (hx : AlgebraicIndependent F x) +include hx + +/-- Canonical isomorphism between rational function field and the + intermediate field generated by algebraically independent elements. -/ +def aevalEquivField : + FractionRing (MvPolynomial ι F) ≃ₐ[F] ↥(IntermediateField.adjoin F (range x)) := + let i := IsFractionRing.liftAlgHom (K := FractionRing (MvPolynomial ι F)) + (algebraicIndependent_iff_injective_aeval.2 hx) + (show _ ≃ₐ[F] i.fieldRange from AlgEquiv.ofInjectiveField i).trans <| + IntermediateField.equivOfEq <| + IsFractionRing.algHom_fieldRange_eq_of_comp_eq_of_range_eq (g := aeval x) (f := i) + (by ext <;> simp [i]) (Algebra.adjoin_range_eq_range_aeval F x).symm + +@[simp] +theorem aevalEquivField_apply_coe (a : FractionRing (MvPolynomial ι F)) : + hx.aevalEquivField a = + IsFractionRing.lift (algebraicIndependent_iff_injective_aeval.2 hx) a := rfl + +theorem aevalEquivField_algebraMap_apply_coe (a : MvPolynomial ι F) : + hx.aevalEquivField (algebraMap _ _ a) = aeval x a := by + simp + +/-- The canonical map from the intermediate field generated by an algebraic independent family + into the rational function field. -/ +def reprField : IntermediateField.adjoin F (range x) →ₐ[F] FractionRing (MvPolynomial ι F) := + hx.aevalEquivField.symm + +@[simp] +theorem lift_reprField (p) : + IsFractionRing.lift (algebraicIndependent_iff_injective_aeval.2 hx) (hx.reprField p) = p := + Subtype.ext_iff.1 (AlgEquiv.apply_symm_apply hx.aevalEquivField p) + +theorem liftAlgHom_comp_reprField : + (IsFractionRing.liftAlgHom (algebraicIndependent_iff_injective_aeval.2 hx)).comp hx.reprField = + IntermediateField.val _ := + AlgHom.ext <| hx.lift_reprField + +end AlgebraicIndependent diff --git a/Mathlib/RingTheory/AlgebraicIndependent/Basic.lean b/Mathlib/RingTheory/AlgebraicIndependent/Basic.lean new file mode 100644 index 0000000000000..605ac77d46b2a --- /dev/null +++ b/Mathlib/RingTheory/AlgebraicIndependent/Basic.lean @@ -0,0 +1,384 @@ +/- +Copyright (c) 2021 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes +-/ +import Mathlib.Algebra.MvPolynomial.Equiv +import Mathlib.Algebra.MvPolynomial.Monad +import Mathlib.Algebra.MvPolynomial.Supported +import Mathlib.RingTheory.AlgebraicIndependent.Defs +import Mathlib.RingTheory.Ideal.Maps +import Mathlib.RingTheory.MvPolynomial.Basic + +/-! +# Algebraic Independence + +This file contains basic results on algebraic independence of a family of elements of an `R`-algebra + +## References + +* [Stacks: Transcendence](https://stacks.math.columbia.edu/tag/030D) + +## TODO +Define the transcendence degree and show it is independent of the choice of a +transcendence basis. + +## Tags +transcendence basis, transcendence degree, transcendence + +-/ + + +noncomputable section + +open Function Set Subalgebra MvPolynomial Algebra + +open scoped Classical + +variable {ι ι' R K A A' : Type*} {x : ι → A} +variable [CommRing R] [CommRing A] [CommRing A'] [Algebra R A] [Algebra R A'] + +theorem algebraicIndependent_iff_ker_eq_bot : + AlgebraicIndependent R x ↔ + RingHom.ker (MvPolynomial.aeval x : MvPolynomial ι R →ₐ[R] A).toRingHom = ⊥ := + RingHom.injective_iff_ker_eq_bot _ + +@[simp] +theorem algebraicIndependent_empty_type_iff [IsEmpty ι] : + AlgebraicIndependent R x ↔ Injective (algebraMap R A) := by + rw [algebraicIndependent_iff_injective_aeval, MvPolynomial.aeval_injective_iff_of_isEmpty] + +namespace AlgebraicIndependent + +variable (hx : AlgebraicIndependent R x) +include hx + +theorem algebraMap_injective : Injective (algebraMap R A) := by + simpa [Function.comp_def] using + (Injective.of_comp_iff (algebraicIndependent_iff_injective_aeval.1 hx) MvPolynomial.C).2 + (MvPolynomial.C_injective _ _) + +theorem linearIndependent : LinearIndependent R x := by + rw [linearIndependent_iff_injective_linearCombination] + have : Finsupp.linearCombination R x = + (MvPolynomial.aeval x).toLinearMap.comp (Finsupp.linearCombination R X) := by + ext + simp + rw [this] + refine (algebraicIndependent_iff_injective_aeval.mp hx).comp ?_ + rw [← linearIndependent_iff_injective_linearCombination] + exact linearIndependent_X _ _ + +protected theorem injective [Nontrivial R] : Injective x := + hx.linearIndependent.injective + +theorem ne_zero [Nontrivial R] (i : ι) : x i ≠ 0 := + hx.linearIndependent.ne_zero i + +theorem map {f : A →ₐ[R] A'} (hf_inj : Set.InjOn f (adjoin R (range x))) : + AlgebraicIndependent R (f ∘ x) := by + have : aeval (f ∘ x) = f.comp (aeval x) := by ext; simp + have h : ∀ p : MvPolynomial ι R, aeval x p ∈ (@aeval R _ _ _ _ _ ((↑) : range x → A)).range := by + intro p + rw [AlgHom.mem_range] + refine ⟨MvPolynomial.rename (codRestrict x (range x) mem_range_self) p, ?_⟩ + simp [Function.comp_def, aeval_rename] + intro x y hxy + rw [this] at hxy + rw [adjoin_eq_range] at hf_inj + exact hx (hf_inj (h x) (h y) hxy) + +theorem map' {f : A →ₐ[R] A'} (hf_inj : Injective f) : AlgebraicIndependent R (f ∘ x) := + hx.map hf_inj.injOn + +/-- If `x = {x_i : A | i : ι}` and `f = {f_i : MvPolynomial ι R | i : ι}` are algebraically +independent over `R`, then `{f_i(x) | i : ι}` is also algebraically independent over `R`. +For the partial converse, see `AlgebraicIndependent.of_aeval`. -/ +theorem aeval_of_algebraicIndependent + {f : ι → MvPolynomial ι R} (hf : AlgebraicIndependent R f) : + AlgebraicIndependent R fun i ↦ aeval x (f i) := by + rw [algebraicIndependent_iff] at hx hf ⊢ + intro p hp + exact hf _ (hx _ (by rwa [← aeval_comp_bind₁, AlgHom.comp_apply] at hp)) + +omit hx in +/-- If `{f_i(x) | i : ι}` is algebraically independent over `R`, then +`{f_i : MvPolynomial ι R | i : ι}` is also algebraically independent over `R`. +In fact, the `x = {x_i : A | i : ι}` is also transcendental over `R` provided that `R` +is a field and `ι` is finite; the proof needs transcendence degree. -/ +theorem of_aeval {f : ι → MvPolynomial ι R} + (H : AlgebraicIndependent R fun i ↦ aeval x (f i)) : + AlgebraicIndependent R f := by + rw [algebraicIndependent_iff] at H ⊢ + intro p hp + exact H p (by rw [← aeval_comp_bind₁, AlgHom.comp_apply, bind₁, hp, map_zero]) + +end AlgebraicIndependent + +theorem MvPolynomial.algebraicIndependent_X (σ R : Type*) [CommRing R] : + AlgebraicIndependent R (X (R := R) (σ := σ)) := by + rw [AlgebraicIndependent, aeval_X_left] + exact injective_id + +open AlgebraicIndependent + +theorem AlgHom.algebraicIndependent_iff (f : A →ₐ[R] A') (hf : Injective f) : + AlgebraicIndependent R (f ∘ x) ↔ AlgebraicIndependent R x := + ⟨fun h => h.of_comp f, fun h => h.map hf.injOn⟩ + +@[nontriviality] +theorem algebraicIndependent_of_subsingleton [Subsingleton R] : AlgebraicIndependent R x := + algebraicIndependent_iff.2 fun _ _ => Subsingleton.elim _ _ + +theorem algebraicIndependent_adjoin (hs : AlgebraicIndependent R x) : + @AlgebraicIndependent ι R (adjoin R (range x)) + (fun i : ι => ⟨x i, subset_adjoin (mem_range_self i)⟩) _ _ _ := + AlgebraicIndependent.of_comp (adjoin R (range x)).val hs + +/-- A set of algebraically independent elements in an algebra `A` over a ring `K` is also +algebraically independent over a subring `R` of `K`. -/ +theorem AlgebraicIndependent.restrictScalars {K : Type*} [CommRing K] [Algebra R K] [Algebra K A] + [IsScalarTower R K A] (hinj : Function.Injective (algebraMap R K)) + (ai : AlgebraicIndependent K x) : AlgebraicIndependent R x := by + have : (aeval x : MvPolynomial ι K →ₐ[K] A).toRingHom.comp (MvPolynomial.map (algebraMap R K)) = + (aeval x : MvPolynomial ι R →ₐ[R] A).toRingHom := by + ext <;> simp [algebraMap_eq_smul_one] + show Injective (aeval x).toRingHom + rw [← this, RingHom.coe_comp] + exact Injective.comp ai (MvPolynomial.map_injective _ hinj) + +section RingHom + +variable {S B FRS FAB : Type*} [CommRing S] [CommRing B] [Algebra S B] + +section + +variable [FunLike FRS R S] [RingHomClass FRS R S] [FunLike FAB A B] [RingHomClass FAB A B] + (f : FRS) (g : FAB) + +theorem AlgebraicIndependent.of_ringHom_of_comp_eq (H : AlgebraicIndependent S (g ∘ x)) + (hf : Function.Injective f) + (h : RingHom.comp (algebraMap S B) f = RingHom.comp g (algebraMap R A)) : + AlgebraicIndependent R x := by + rw [algebraicIndependent_iff] at H ⊢ + intro p hp + have := H (p.map f) <| by + have : (g : A →+* B) _ = _ := congr(g $hp) + rwa [map_zero, map_aeval, ← h, ← eval₂Hom_map_hom, ← aeval_eq_eval₂Hom] at this + exact map_injective (f : R →+* S) hf (by rwa [map_zero]) + +theorem AlgebraicIndependent.ringHom_of_comp_eq (H : AlgebraicIndependent R x) + (hf : Function.Surjective f) (hg : Function.Injective g) + (h : RingHom.comp (algebraMap S B) f = RingHom.comp g (algebraMap R A)) : + AlgebraicIndependent S (g ∘ x) := by + rw [algebraicIndependent_iff] at H ⊢ + intro p hp + obtain ⟨q, rfl⟩ := map_surjective (f : R →+* S) hf p + rw [H q (hg (by rwa [map_zero, ← RingHom.coe_coe g, map_aeval, ← h, ← eval₂Hom_map_hom, + ← aeval_eq_eval₂Hom])), map_zero] + +end + +section + +variable [EquivLike FRS R S] [RingEquivClass FRS R S] [FunLike FAB A B] [RingHomClass FAB A B] + (f : FRS) (g : FAB) + +theorem algebraicIndependent_ringHom_iff_of_comp_eq + (hg : Function.Injective g) + (h : RingHom.comp (algebraMap S B) f = RingHom.comp g (algebraMap R A)) : + AlgebraicIndependent S (g ∘ x) ↔ AlgebraicIndependent R x := + ⟨fun H ↦ H.of_ringHom_of_comp_eq f g (EquivLike.injective f) h, + fun H ↦ H.ringHom_of_comp_eq f g (EquivLike.surjective f) hg h⟩ + +end + +end RingHom + +/-- Every finite subset of an algebraically independent set is algebraically independent. -/ +theorem algebraicIndependent_finset_map_embedding_subtype (s : Set A) + (li : AlgebraicIndependent R ((↑) : s → A)) (t : Finset s) : + AlgebraicIndependent R ((↑) : Finset.map (Embedding.subtype s) t → A) := by + let f : t.map (Embedding.subtype s) → s := fun x => + ⟨x.1, by + obtain ⟨x, h⟩ := x + rw [Finset.mem_map] at h + obtain ⟨a, _, rfl⟩ := h + simp only [Subtype.coe_prop, Embedding.coe_subtype]⟩ + convert AlgebraicIndependent.comp li f _ + rintro ⟨x, hx⟩ ⟨y, hy⟩ + rw [Finset.mem_map] at hx hy + obtain ⟨a, _, rfl⟩ := hx + obtain ⟨b, _, rfl⟩ := hy + simp only [f, imp_self, Subtype.mk_eq_mk] + +/-- If every finite set of algebraically independent element has cardinality at most `n`, +then the same is true for arbitrary sets of algebraically independent elements. -/ +theorem algebraicIndependent_bounded_of_finset_algebraicIndependent_bounded {n : ℕ} + (H : ∀ s : Finset A, (AlgebraicIndependent R fun i : s => (i : A)) → s.card ≤ n) : + ∀ s : Set A, AlgebraicIndependent R ((↑) : s → A) → Cardinal.mk s ≤ n := by + intro s li + apply Cardinal.card_le_of + intro t + rw [← Finset.card_map (Embedding.subtype s)] + apply H + apply algebraicIndependent_finset_map_embedding_subtype _ li + +section Subtype + +theorem AlgebraicIndependent.restrict_of_comp_subtype {s : Set ι} + (hs : AlgebraicIndependent R (x ∘ (↑) : s → A)) : AlgebraicIndependent R (s.restrict x) := + hs + +variable (R A) + +theorem algebraicIndependent_empty_iff : + AlgebraicIndependent R ((↑) : (∅ : Set A) → A) ↔ Injective (algebraMap R A) := by simp + +end Subtype + +theorem AlgebraicIndependent.to_subtype_range {ι} {f : ι → A} (hf : AlgebraicIndependent R f) : + AlgebraicIndependent R ((↑) : range f → A) := by + nontriviality R + rwa [algebraicIndependent_subtype_range hf.injective] + +theorem AlgebraicIndependent.to_subtype_range' {ι} {f : ι → A} (hf : AlgebraicIndependent R f) {t} + (ht : range f = t) : AlgebraicIndependent R ((↑) : t → A) := + ht ▸ hf.to_subtype_range + +theorem algebraicIndependent_comp_subtype {s : Set ι} : + AlgebraicIndependent R (x ∘ (↑) : s → A) ↔ + ∀ p ∈ MvPolynomial.supported R s, aeval x p = 0 → p = 0 := by + have : (aeval (x ∘ (↑) : s → A) : _ →ₐ[R] _) = (aeval x).comp (rename (↑)) := by ext; simp + have : ∀ p : MvPolynomial s R, rename ((↑) : s → ι) p = 0 ↔ p = 0 := + (injective_iff_map_eq_zero' (rename ((↑) : s → ι) : MvPolynomial s R →ₐ[R] _).toRingHom).1 + (rename_injective _ Subtype.val_injective) + simp [algebraicIndependent_iff, supported_eq_range_rename, *] + +theorem algebraicIndependent_subtype {s : Set A} : + AlgebraicIndependent R ((↑) : s → A) ↔ + ∀ p : MvPolynomial A R, p ∈ MvPolynomial.supported R s → aeval id p = 0 → p = 0 := by + apply @algebraicIndependent_comp_subtype _ _ _ id + +theorem algebraicIndependent_of_finite (s : Set A) + (H : ∀ t ⊆ s, t.Finite → AlgebraicIndependent R ((↑) : t → A)) : + AlgebraicIndependent R ((↑) : s → A) := + algebraicIndependent_subtype.2 fun p hp => + algebraicIndependent_subtype.1 (H _ (mem_supported.1 hp) (Finset.finite_toSet _)) _ (by simp) + +theorem AlgebraicIndependent.image_of_comp {ι ι'} (s : Set ι) (f : ι → ι') (g : ι' → A) + (hs : AlgebraicIndependent R fun x : s => g (f x)) : + AlgebraicIndependent R fun x : f '' s => g x := by + nontriviality R + have : InjOn f s := injOn_iff_injective.2 hs.injective.of_comp + exact (algebraicIndependent_equiv' (Equiv.Set.imageOfInjOn f s this) rfl).1 hs + +theorem AlgebraicIndependent.image {ι} {s : Set ι} {f : ι → A} + (hs : AlgebraicIndependent R fun x : s => f x) : + AlgebraicIndependent R fun x : f '' s => (x : A) := by + convert AlgebraicIndependent.image_of_comp s f id hs + +theorem algebraicIndependent_iUnion_of_directed {η : Type*} [Nonempty η] {s : η → Set A} + (hs : Directed (· ⊆ ·) s) (h : ∀ i, AlgebraicIndependent R ((↑) : s i → A)) : + AlgebraicIndependent R ((↑) : (⋃ i, s i) → A) := by + refine algebraicIndependent_of_finite (⋃ i, s i) fun t ht ft => ?_ + rcases finite_subset_iUnion ft ht with ⟨I, fi, hI⟩ + rcases hs.finset_le fi.toFinset with ⟨i, hi⟩ + exact (h i).mono (Subset.trans hI <| iUnion₂_subset fun j hj => hi j (fi.mem_toFinset.2 hj)) + +theorem algebraicIndependent_sUnion_of_directed {s : Set (Set A)} (hsn : s.Nonempty) + (hs : DirectedOn (· ⊆ ·) s) (h : ∀ a ∈ s, AlgebraicIndependent R ((↑) : a → A)) : + AlgebraicIndependent R ((↑) : ⋃₀ s → A) := by + letI : Nonempty s := Nonempty.to_subtype hsn + rw [sUnion_eq_iUnion] + exact algebraicIndependent_iUnion_of_directed hs.directed_val (by simpa using h) + +theorem exists_maximal_algebraicIndependent (s t : Set A) (hst : s ⊆ t) + (hs : AlgebraicIndependent R ((↑) : s → A)) : ∃ u, s ⊆ u ∧ + Maximal (fun (x : Set A) ↦ AlgebraicIndependent R ((↑) : x → A) ∧ x ⊆ t) u := by + refine zorn_subset_nonempty { u : Set A | AlgebraicIndependent R ((↑) : u → A) ∧ u ⊆ t} + (fun c hc chainc hcn ↦ ⟨⋃₀ c, ⟨?_, ?_⟩, fun _ ↦ subset_sUnion_of_mem⟩) s ⟨hs, hst⟩ + · exact algebraicIndependent_sUnion_of_directed hcn chainc.directedOn (fun x hxc ↦ (hc hxc).1) + exact fun x ⟨w, hyc, hwy⟩ ↦ (hc hyc).2 hwy + +theorem AlgebraicIndependent.repr_ker (hx : AlgebraicIndependent R x) : + RingHom.ker (hx.repr : adjoin R (range x) →+* MvPolynomial ι R) = ⊥ := + (RingHom.injective_iff_ker_eq_bot _).1 (AlgEquiv.injective _) + +-- TODO - make this an `AlgEquiv` +/-- The isomorphism between `MvPolynomial (Option ι) R` and the polynomial ring over +the algebra generated by an algebraically independent family. -/ +def AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin (hx : AlgebraicIndependent R x) : + MvPolynomial (Option ι) R ≃+* Polynomial (adjoin R (Set.range x)) := + (MvPolynomial.optionEquivLeft _ _).toRingEquiv.trans + (Polynomial.mapEquiv hx.aevalEquiv.toRingEquiv) + +@[simp] +theorem AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_apply + (hx : AlgebraicIndependent R x) (y) : + hx.mvPolynomialOptionEquivPolynomialAdjoin y = + Polynomial.map (hx.aevalEquiv : MvPolynomial ι R →+* adjoin R (range x)) + (aeval (fun o : Option ι => o.elim Polynomial.X fun s : ι => Polynomial.C (X s)) y) := + rfl + +/-- `simp`-normal form of `mvPolynomialOptionEquivPolynomialAdjoin_C` -/ +@[simp] +theorem AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_C' + (hx : AlgebraicIndependent R x) (r) : + Polynomial.C (hx.aevalEquiv (C r)) = Polynomial.C (algebraMap _ _ r) := by + congr + apply_fun Subtype.val using Subtype.val_injective + simp + +theorem AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_C + (hx : AlgebraicIndependent R x) (r) : + hx.mvPolynomialOptionEquivPolynomialAdjoin (C r) = Polynomial.C (algebraMap _ _ r) := by + simp + +theorem AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_X_none + (hx : AlgebraicIndependent R x) : + hx.mvPolynomialOptionEquivPolynomialAdjoin (X none) = Polynomial.X := by + rw [AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_apply, aeval_X, Option.elim, + Polynomial.map_X] + +theorem AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_X_some + (hx : AlgebraicIndependent R x) (i) : + hx.mvPolynomialOptionEquivPolynomialAdjoin (X (some i)) = + Polynomial.C (hx.aevalEquiv (X i)) := by + rw [AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_apply, aeval_X, Option.elim, + Polynomial.map_C, RingHom.coe_coe] + +theorem AlgebraicIndependent.aeval_comp_mvPolynomialOptionEquivPolynomialAdjoin + (hx : AlgebraicIndependent R x) (a : A) : + RingHom.comp + (↑(Polynomial.aeval a : Polynomial (adjoin R (Set.range x)) →ₐ[_] A) : + Polynomial (adjoin R (Set.range x)) →+* A) + hx.mvPolynomialOptionEquivPolynomialAdjoin.toRingHom = + ↑(MvPolynomial.aeval fun o : Option ι => o.elim a x : MvPolynomial (Option ι) R →ₐ[R] A) := by + refine MvPolynomial.ringHom_ext ?_ ?_ <;> + simp only [RingHom.comp_apply, RingEquiv.toRingHom_eq_coe, RingEquiv.coe_toRingHom, + AlgHom.coe_toRingHom, AlgHom.coe_toRingHom] + · intro r + rw [hx.mvPolynomialOptionEquivPolynomialAdjoin_C, aeval_C, Polynomial.aeval_C, + IsScalarTower.algebraMap_apply R (adjoin R (range x)) A] + · rintro (⟨⟩ | ⟨i⟩) + · rw [hx.mvPolynomialOptionEquivPolynomialAdjoin_X_none, aeval_X, Polynomial.aeval_X, + Option.elim] + · rw [hx.mvPolynomialOptionEquivPolynomialAdjoin_X_some, Polynomial.aeval_C, + hx.algebraMap_aevalEquiv, aeval_X, aeval_X, Option.elim] + +section Field + +variable [Field K] [Algebra K A] + +/- Porting note: removing `simp`, not in simp normal form. Could make `Function.Injective f` a +simp lemma when `f` is a field hom, and then simp would prove this -/ +theorem algebraicIndependent_empty_type [IsEmpty ι] [Nontrivial A] : AlgebraicIndependent K x := by + rw [algebraicIndependent_empty_type_iff] + exact RingHom.injective _ + +theorem algebraicIndependent_empty [Nontrivial A] : + AlgebraicIndependent K ((↑) : (∅ : Set A) → A) := + algebraicIndependent_empty_type + +end Field diff --git a/Mathlib/RingTheory/AlgebraicIndependent/Defs.lean b/Mathlib/RingTheory/AlgebraicIndependent/Defs.lean new file mode 100644 index 0000000000000..c3d9b6df75314 --- /dev/null +++ b/Mathlib/RingTheory/AlgebraicIndependent/Defs.lean @@ -0,0 +1,155 @@ +/- +Copyright (c) 2021 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes +-/ +import Mathlib.Algebra.MvPolynomial.CommRing + +/-! +# Algebraic Independence + +This file defines algebraic independence of a family of elements of an `R` algebra. + +## Main definitions + +* `AlgebraicIndependent` - `AlgebraicIndependent R x` states the family of elements `x` + is algebraically independent over `R`, meaning that the canonical map out of the multivariable + polynomial ring is injective. + +* `AlgebraicIndependent.aevalEquiv` - The canonical isomorphism from the polynomial ring to the + subalgebra generated by an algebraic independent family. + +* `AlgebraicIndependent.repr` - The canonical map from the subalgebra generated by an + algebraic independent family into the polynomial ring. It is the inverse of + `AlgebraicIndependent.aevalEquiv`. + +* `IsTranscendenceBasis R x` - a family `x` is a transcendence basis over `R` if it is a maximal + algebraically independent subset. + +## Main results + +We show that algebraic independence is preserved under injective maps of the indices. + +## References + +* [Stacks: Transcendence](https://stacks.math.columbia.edu/tag/030D) + +-/ + + +noncomputable section + +open Function Set Subalgebra MvPolynomial Algebra + +open scoped Classical + +variable {ι ι' : Type*} (R : Type*) {K A A' : Type*} (x : ι → A) +variable [CommRing R] [CommRing A] [CommRing A'] [Algebra R A] [Algebra R A'] + +/-- `AlgebraicIndependent R x` states the family of elements `x` + is algebraically independent over `R`, meaning that the canonical + map out of the multivariable polynomial ring is injective. -/ +def AlgebraicIndependent : Prop := + Injective (MvPolynomial.aeval x : MvPolynomial ι R →ₐ[R] A) + +variable {R} {x} + +theorem algebraicIndependent_iff : + AlgebraicIndependent R x ↔ + ∀ p : MvPolynomial ι R, MvPolynomial.aeval (x : ι → A) p = 0 → p = 0 := + injective_iff_map_eq_zero _ + +theorem AlgebraicIndependent.eq_zero_of_aeval_eq_zero (h : AlgebraicIndependent R x) : + ∀ p : MvPolynomial ι R, MvPolynomial.aeval (x : ι → A) p = 0 → p = 0 := + algebraicIndependent_iff.1 h + +theorem algebraicIndependent_iff_injective_aeval : + AlgebraicIndependent R x ↔ Injective (MvPolynomial.aeval x : MvPolynomial ι R →ₐ[R] A) := + Iff.rfl + +namespace AlgebraicIndependent + +theorem of_comp (f : A →ₐ[R] A') (hfv : AlgebraicIndependent R (f ∘ x)) : + AlgebraicIndependent R x := by + have : aeval (f ∘ x) = f.comp (aeval x) := by ext; simp + rw [AlgebraicIndependent, this, AlgHom.coe_comp] at hfv + exact hfv.of_comp + +variable (hx : AlgebraicIndependent R x) +include hx + +theorem comp (f : ι' → ι) (hf : Function.Injective f) : AlgebraicIndependent R (x ∘ f) := by + intro p q + simpa [aeval_rename, (rename_injective f hf).eq_iff] using @hx (rename f p) (rename f q) + +theorem coe_range : AlgebraicIndependent R ((↑) : range x → A) := by + simpa using hx.comp _ (rangeSplitting_injective x) + +end AlgebraicIndependent + +open AlgebraicIndependent + +theorem algebraicIndependent_equiv (e : ι ≃ ι') {f : ι' → A} : + AlgebraicIndependent R (f ∘ e) ↔ AlgebraicIndependent R f := + ⟨fun h => Function.comp_id f ▸ e.self_comp_symm ▸ h.comp _ e.symm.injective, + fun h => h.comp _ e.injective⟩ + +theorem algebraicIndependent_equiv' (e : ι ≃ ι') {f : ι' → A} {g : ι → A} (h : f ∘ e = g) : + AlgebraicIndependent R g ↔ AlgebraicIndependent R f := + h ▸ algebraicIndependent_equiv e + +theorem algebraicIndependent_subtype_range {ι} {f : ι → A} (hf : Injective f) : + AlgebraicIndependent R ((↑) : range f → A) ↔ AlgebraicIndependent R f := + Iff.symm <| algebraicIndependent_equiv' (Equiv.ofInjective f hf) rfl + +alias ⟨AlgebraicIndependent.of_subtype_range, _⟩ := algebraicIndependent_subtype_range + +theorem algebraicIndependent_image {ι} {s : Set ι} {f : ι → A} (hf : Set.InjOn f s) : + (AlgebraicIndependent R fun x : s => f x) ↔ AlgebraicIndependent R fun x : f '' s => (x : A) := + algebraicIndependent_equiv' (Equiv.Set.imageOfInjOn _ _ hf) rfl + +namespace AlgebraicIndependent + +theorem mono {t s : Set A} (h : t ⊆ s) + (hx : AlgebraicIndependent R ((↑) : s → A)) : AlgebraicIndependent R ((↑) : t → A) := by + simpa [Function.comp] using hx.comp (inclusion h) (inclusion_injective h) + +section repr + +variable (hx : AlgebraicIndependent R x) +include hx + +/-- Canonical isomorphism between polynomials and the subalgebra generated by + algebraically independent elements. -/ +@[simps! apply_coe] +def aevalEquiv : MvPolynomial ι R ≃ₐ[R] Algebra.adjoin R (range x) := + (AlgEquiv.ofInjective (aeval x) (algebraicIndependent_iff_injective_aeval.1 hx)).trans + (Subalgebra.equivOfEq _ _ (Algebra.adjoin_range_eq_range_aeval R x).symm) + +--@[simp] Porting note: removing simp because the linter complains about deterministic timeout +theorem algebraMap_aevalEquiv (p : MvPolynomial ι R) : + algebraMap (Algebra.adjoin R (range x)) A (hx.aevalEquiv p) = aeval x p := + rfl + +/-- The canonical map from the subalgebra generated by an algebraic independent family + into the polynomial ring. -/ +def repr : Algebra.adjoin R (range x) →ₐ[R] MvPolynomial ι R := + hx.aevalEquiv.symm + +@[simp] +theorem aeval_repr (p) : aeval x (hx.repr p) = p := + Subtype.ext_iff.1 (AlgEquiv.apply_symm_apply hx.aevalEquiv p) + +theorem aeval_comp_repr : (aeval x).comp hx.repr = Subalgebra.val _ := + AlgHom.ext hx.aeval_repr + +end repr + +end AlgebraicIndependent + +variable (R) + +/-- A family is a transcendence basis if it is a maximal algebraically independent subset. -/ +def IsTranscendenceBasis (x : ι → A) : Prop := + AlgebraicIndependent R x ∧ + ∀ (s : Set A) (_ : AlgebraicIndependent R ((↑) : s → A)) (_ : range x ≤ s), range x = s diff --git a/Mathlib/RingTheory/AlgebraicIndependent/RankAndCardinality.lean b/Mathlib/RingTheory/AlgebraicIndependent/RankAndCardinality.lean new file mode 100644 index 0000000000000..da5e72bcfd342 --- /dev/null +++ b/Mathlib/RingTheory/AlgebraicIndependent/RankAndCardinality.lean @@ -0,0 +1,93 @@ +/- +Copyright (c) 2021 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes +-/ +import Mathlib.FieldTheory.MvRatFunc.Rank +import Mathlib.RingTheory.Algebraic.Cardinality +import Mathlib.RingTheory.AlgebraicIndependent.Adjoin +import Mathlib.RingTheory.AlgebraicIndependent.Transcendental +import Mathlib.RingTheory.AlgebraicIndependent.TranscendenceBasis + +/-! +# Cardinality of a transcendence basis + +This file concerns the cardinality of a transcendence basis. + +## References + +* [Stacks: Transcendence](https://stacks.math.columbia.edu/tag/030D) + +## TODO +Define the transcendence degree and show it is independent of the choice of a +transcendence basis. + +## Tags +transcendence basis, transcendence degree, transcendence + +-/ + +noncomputable section + +open Function Set Subalgebra MvPolynomial Algebra + +open scoped Classical + +universe u v w + +open AlgebraicIndependent + +open Cardinal + +theorem IsTranscendenceBasis.lift_cardinalMk_eq_max_lift + {F : Type u} {E : Type v} [CommRing F] [Nontrivial F] [CommRing E] [IsDomain E] [Algebra F E] + {ι : Type w} {x : ι → E} [Nonempty ι] (hx : IsTranscendenceBasis F x) : + lift.{max u w} #E = lift.{max v w} #F ⊔ lift.{max u v} #ι ⊔ ℵ₀ := by + let K := Algebra.adjoin F (Set.range x) + suffices #E = #K by simp [this, ← lift_mk_eq'.2 ⟨hx.1.aevalEquiv.toEquiv⟩] + haveI : Algebra.IsAlgebraic K E := hx.isAlgebraic + refine le_antisymm ?_ (mk_le_of_injective Subtype.val_injective) + haveI : Infinite K := hx.1.aevalEquiv.infinite_iff.1 inferInstance + simpa only [sup_eq_left.2 (aleph0_le_mk K)] using Algebra.IsAlgebraic.cardinalMk_le_max K E + +theorem IsTranscendenceBasis.lift_rank_eq_max_lift + {F : Type u} {E : Type v} [Field F] [Field E] [Algebra F E] + {ι : Type w} {x : ι → E} [Nonempty ι] (hx : IsTranscendenceBasis F x) : + lift.{max u w} (Module.rank F E) = lift.{max v w} #F ⊔ lift.{max u v} #ι ⊔ ℵ₀ := by + let K := IntermediateField.adjoin F (Set.range x) + haveI : Algebra.IsAlgebraic K E := hx.isAlgebraic_field + rw [← rank_mul_rank F K E, lift_mul, ← hx.1.aevalEquivField.toLinearEquiv.lift_rank_eq, + MvRatFunc.rank_eq_max_lift, lift_max, lift_max, lift_lift, lift_lift, lift_aleph0] + refine mul_eq_left le_sup_right ((lift_le.2 ((rank_le_card K E).trans + (Algebra.IsAlgebraic.cardinalMk_le_max K E))).trans_eq ?_) (by simp [rank_pos.ne']) + simp [← lift_mk_eq'.2 ⟨hx.1.aevalEquivField.toEquiv⟩] + +theorem Algebra.Transcendental.rank_eq_cardinalMk + (F : Type u) (E : Type v) [Field F] [Field E] [Algebra F E] [Algebra.Transcendental F E] : + Module.rank F E = #E := by + obtain ⟨ι, x, hx⟩ := exists_isTranscendenceBasis' _ (algebraMap F E).injective + haveI := hx.nonempty_iff_transcendental.2 ‹_› + simpa [← hx.lift_cardinalMk_eq_max_lift] using hx.lift_rank_eq_max_lift + +theorem IntermediateField.rank_sup_le + {F : Type u} {E : Type v} [Field F] [Field E] [Algebra F E] (A B : IntermediateField F E) : + Module.rank F ↥(A ⊔ B) ≤ Module.rank F A * Module.rank F B := by + by_cases hA : Algebra.IsAlgebraic F A + · exact rank_sup_le_of_isAlgebraic A B (Or.inl hA) + by_cases hB : Algebra.IsAlgebraic F B + · exact rank_sup_le_of_isAlgebraic A B (Or.inr hB) + rw [← Algebra.transcendental_iff_not_isAlgebraic] at hA hB + haveI : Algebra.Transcendental F ↥(A ⊔ B) := .ringHom_of_comp_eq (RingHom.id F) + (inclusion le_sup_left) Function.surjective_id (inclusion_injective _) rfl + haveI := Algebra.Transcendental.infinite F A + haveI := Algebra.Transcendental.infinite F B + simp_rw [Algebra.Transcendental.rank_eq_cardinalMk] + rw [sup_def, mul_mk_eq_max, ← Cardinal.lift_le.{u}] + refine (lift_cardinalMk_adjoin_le _ _).trans ?_ + calc + _ ≤ Cardinal.lift.{v} #F ⊔ Cardinal.lift.{u} (#A ⊔ #B) ⊔ ℵ₀ := by + gcongr + rw [Cardinal.lift_le] + exact (mk_union_le _ _).trans_eq (by simp) + _ = _ := by + simp [lift_mk_le_lift_mk_of_injective (algebraMap F A).injective] diff --git a/Mathlib/RingTheory/AlgebraicIndependent/TranscendenceBasis.lean b/Mathlib/RingTheory/AlgebraicIndependent/TranscendenceBasis.lean new file mode 100644 index 0000000000000..db169e8b20d4c --- /dev/null +++ b/Mathlib/RingTheory/AlgebraicIndependent/TranscendenceBasis.lean @@ -0,0 +1,129 @@ +/- +Copyright (c) 2021 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes +-/ +import Mathlib.FieldTheory.Adjoin +import Mathlib.RingTheory.AlgebraicIndependent.Transcendental + +/-! +# Transcendence basis + +This file defines the transcendence basis as a maximal algebraically independent subset. + +## Main results + +* `exists_isTranscendenceBasis`: a ring extension has a transcendence basis + +## References + +* [Stacks: Transcendence](https://stacks.math.columbia.edu/tag/030D) + +## TODO +Define the transcendence degree and show it is independent of the choice of a +transcendence basis. + +## Tags +transcendence basis, transcendence degree, transcendence + +-/ + +noncomputable section + +open Function Set Subalgebra MvPolynomial Algebra + +open scoped Classical + +universe u v w + +variable {ι ι' : Type*} (R : Type*) {K A A' : Type*} +variable {x : ι → A} +variable [CommRing R] [CommRing A] [CommRing A'] [Algebra R A] [Algebra R A'] + +open AlgebraicIndependent + +theorem exists_isTranscendenceBasis (h : Injective (algebraMap R A)) : + ∃ s : Set A, IsTranscendenceBasis R ((↑) : s → A) := by + cases' exists_maximal_algebraicIndependent (∅ : Set A) Set.univ (Set.subset_univ _) + ((algebraicIndependent_empty_iff R A).2 h) with + s hs + refine ⟨s, hs.2.1.1, fun t ht hst ↦ ?_⟩ + simp only [Subtype.range_coe_subtype, setOf_mem_eq] at * + exact hs.2.eq_of_le ⟨ht, subset_univ _⟩ hst + +/-- `Type` version of `exists_isTranscendenceBasis`. -/ +theorem exists_isTranscendenceBasis' (R : Type u) {A : Type v} [CommRing R] [CommRing A] + [Algebra R A] (h : Injective (algebraMap R A)) : + ∃ (ι : Type v) (x : ι → A), IsTranscendenceBasis R x := by + obtain ⟨s, h⟩ := exists_isTranscendenceBasis R h + exact ⟨s, Subtype.val, h⟩ + +variable {R} + +theorem AlgebraicIndependent.isTranscendenceBasis_iff {ι : Type w} {R : Type u} [CommRing R] + [Nontrivial R] {A : Type v} [CommRing A] [Algebra R A] {x : ι → A} + (i : AlgebraicIndependent R x) : + IsTranscendenceBasis R x ↔ + ∀ (κ : Type v) (w : κ → A) (_ : AlgebraicIndependent R w) (j : ι → κ) (_ : w ∘ j = x), + Surjective j := by + fconstructor + · rintro p κ w i' j rfl + have p := p.2 (range w) i'.coe_range (range_comp_subset_range _ _) + rw [range_comp, ← @image_univ _ _ w] at p + exact range_eq_univ.mp (image_injective.mpr i'.injective p) + · intro p + use i + intro w i' h + specialize p w ((↑) : w → A) i' (fun i => ⟨x i, range_subset_iff.mp h i⟩) (by ext; simp) + have q := congr_arg (fun s => ((↑) : w → A) '' s) p.range_eq + dsimp at q + rw [← image_univ, image_image] at q + simpa using q + +theorem IsTranscendenceBasis.isAlgebraic [Nontrivial R] (hx : IsTranscendenceBasis R x) : + Algebra.IsAlgebraic (adjoin R (range x)) A := by + constructor + intro a + rw [← not_iff_comm.1 (hx.1.option_iff _).symm] + intro ai + have h₁ : range x ⊆ range fun o : Option ι => o.elim a x := by + rintro x ⟨y, rfl⟩ + exact ⟨some y, rfl⟩ + have h₂ : range x ≠ range fun o : Option ι => o.elim a x := by + intro h + have : a ∈ range x := by + rw [h] + exact ⟨none, rfl⟩ + rcases this with ⟨b, rfl⟩ + have : some b = none := ai.injective rfl + simpa + exact h₂ (hx.2 (Set.range fun o : Option ι => o.elim a x) + ((algebraicIndependent_subtype_range ai.injective).2 ai) h₁) + +/-- If `x` is a transcendence basis of `A/R`, then it is empty if and only if +`A/R` is algebraic. -/ +theorem IsTranscendenceBasis.isEmpty_iff_isAlgebraic [Nontrivial R] + (hx : IsTranscendenceBasis R x) : + IsEmpty ι ↔ Algebra.IsAlgebraic R A := by + refine ⟨fun _ ↦ ?_, fun _ ↦ hx.1.isEmpty_of_isAlgebraic⟩ + have := hx.isAlgebraic + rw [Set.range_eq_empty x, adjoin_empty] at this + exact algebra_isAlgebraic_of_algebra_isAlgebraic_bot_left R A + +/-- If `x` is a transcendence basis of `A/R`, then it is not empty if and only if +`A/R` is transcendental. -/ +theorem IsTranscendenceBasis.nonempty_iff_transcendental [Nontrivial R] + (hx : IsTranscendenceBasis R x) : + Nonempty ι ↔ Algebra.Transcendental R A := by + rw [← not_isEmpty_iff, Algebra.transcendental_iff_not_isAlgebraic, hx.isEmpty_iff_isAlgebraic] + +theorem IsTranscendenceBasis.isAlgebraic_field {F E : Type*} {x : ι → E} + [Field F] [Field E] [Algebra F E] (hx : IsTranscendenceBasis F x) : + Algebra.IsAlgebraic (IntermediateField.adjoin F (range x)) E := by + haveI := hx.isAlgebraic + set S := range x + letI : Algebra (adjoin F S) (IntermediateField.adjoin F S) := + (Subalgebra.inclusion (IntermediateField.algebra_adjoin_le_adjoin F S)).toRingHom.toAlgebra + haveI : IsScalarTower (adjoin F S) (IntermediateField.adjoin F S) E := + IsScalarTower.of_algebraMap_eq (congrFun rfl) + exact Algebra.IsAlgebraic.extendScalars (R := adjoin F S) (Subalgebra.inclusion_injective _) diff --git a/Mathlib/RingTheory/AlgebraicIndependent/Transcendental.lean b/Mathlib/RingTheory/AlgebraicIndependent/Transcendental.lean new file mode 100644 index 0000000000000..332493c379080 --- /dev/null +++ b/Mathlib/RingTheory/AlgebraicIndependent/Transcendental.lean @@ -0,0 +1,146 @@ +/- +Copyright (c) 2021 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes +-/ +import Mathlib.Data.Fin.Tuple.Reflection +import Mathlib.RingTheory.Algebraic.MvPolynomial +import Mathlib.RingTheory.AlgebraicIndependent.Basic + +/-! +# Algebraic Independence + +This file relates algebraic independence and transcendence (or algebraicity) of elements. + +## References + +* [Stacks: Transcendence](https://stacks.math.columbia.edu/tag/030D) + +## Tags +transcendence + +-/ + +noncomputable section + +open Function Set Subalgebra MvPolynomial Algebra + +open scoped Classical + +variable {ι ι' R K A A' : Type*} {x : ι → A} +variable [CommRing R] [CommRing A] [CommRing A'] [Algebra R A] [Algebra R A'] + +/-- A one-element family `x` is algebraically independent if and only if +its element is transcendental. -/ +@[simp] +theorem algebraicIndependent_unique_type_iff [Unique ι] : + AlgebraicIndependent R x ↔ Transcendental R (x default) := by + rw [transcendental_iff_injective, algebraicIndependent_iff_injective_aeval] + let i := (renameEquiv R (Equiv.equivPUnit.{_, 1} ι)).trans (pUnitAlgEquiv R) + have key : aeval (R := R) x = (Polynomial.aeval (R := R) (x default)).comp i := by + ext y + simp [i, Subsingleton.elim y default] + simp [key] + +/-- The one-element family `![x]` is algebraically independent if and only if +`x` is transcendental. -/ +theorem algebraicIndependent_iff_transcendental {x : A} : + AlgebraicIndependent R ![x] ↔ Transcendental R x := by + simp + +namespace AlgebraicIndependent + +variable (hx : AlgebraicIndependent R x) +include hx + +/-- If a family `x` is algebraically independent, then any of its element is transcendental. -/ +theorem transcendental (i : ι) : Transcendental R (x i) := by + have := hx.comp ![i] (Function.injective_of_subsingleton _) + have : AlgebraicIndependent R ![x i] := by rwa [← FinVec.map_eq] at this + rwa [← algebraicIndependent_iff_transcendental] + +/-- If `A/R` is algebraic, then all algebraically independent families are empty. -/ +theorem isEmpty_of_isAlgebraic [Algebra.IsAlgebraic R A] : IsEmpty ι := by + rcases isEmpty_or_nonempty ι with h | ⟨⟨i⟩⟩ + · exact h + exact False.elim (hx.transcendental i (Algebra.IsAlgebraic.isAlgebraic _)) + +end AlgebraicIndependent + +open AlgebraicIndependent + +theorem AlgebraicIndependent.option_iff (hx : AlgebraicIndependent R x) (a : A) : + (AlgebraicIndependent R fun o : Option ι => o.elim a x) ↔ + Transcendental (adjoin R (Set.range x)) a := by + rw [algebraicIndependent_iff_injective_aeval, transcendental_iff_injective, + ← AlgHom.coe_toRingHom, ← hx.aeval_comp_mvPolynomialOptionEquivPolynomialAdjoin, + RingHom.coe_comp] + exact Injective.of_comp_iff' (Polynomial.aeval a) + (mvPolynomialOptionEquivPolynomialAdjoin hx).bijective + +/-- Variant of `algebraicIndependent_of_finite` using `Transcendental`. -/ +theorem algebraicIndependent_of_finite' (s : Set A) + (hinj : Injective (algebraMap R A)) + (H : ∀ t ⊆ s, t.Finite → ∀ a ∈ s, a ∉ t → Transcendental (adjoin R t) a) : + AlgebraicIndependent R ((↑) : s → A) := by + classical + refine algebraicIndependent_of_finite s fun t hts hfin ↦ hfin.induction_on' + ((algebraicIndependent_empty_iff R A).2 hinj) fun {a} {u} ha hu ha' h ↦ ?_ + convert ((Subtype.range_coe ▸ h.option_iff a).2 <| H u (hu.trans hts) (hfin.subset hu) + a (hts ha) ha').comp _ (Set.subtypeInsertEquivOption ha').injective + ext x + by_cases h : ↑x = a <;> simp [h, Set.subtypeInsertEquivOption] + +/-- `Type` version of `algebraicIndependent_of_finite'`. -/ +theorem algebraicIndependent_of_finite_type' + (hinj : Injective (algebraMap R A)) + (H : ∀ t : Set ι, t.Finite → ∀ i : ι, i ∉ t → Transcendental (adjoin R (x '' t)) (x i)) : + AlgebraicIndependent R x := by + nontriviality R + haveI := hinj.nontrivial + have hx : Injective x := by + simp_rw [Transcendental] at H + contrapose! H + obtain ⟨i, j, h1, h2⟩ := not_injective_iff.1 H + refine ⟨{j}, by simp, i, by simp [h2], ?_⟩ + rw [h1, Set.image_singleton] + exact isAlgebraic_algebraMap (⟨x j, Algebra.self_mem_adjoin_singleton R _⟩ : adjoin R {x j}) + rw [← Set.coe_comp_rangeFactorization x] + refine .comp (algebraicIndependent_of_finite' _ hinj fun t ht hfin a ha ha' ↦ ?_) _ + (Set.rightInverse_rangeSplitting hx).injective + change Finite t at hfin + have := H (x ⁻¹' t) (Finite.of_injective _ (hx.restrictPreimage t)) + ((Equiv.ofInjective _ hx).symm ⟨_, ha⟩) + (by rwa [Set.mem_preimage, Equiv.apply_ofInjective_symm hx]) + rwa [Set.image_preimage_eq_inter_range, Set.inter_eq_self_of_subset_left ht, + Equiv.apply_ofInjective_symm hx] at this + +namespace MvPolynomial + +/-- If for each `i : ι`, `f_i : R[X]` is transcendental over `R`, then `{f_i(X_i) | i : ι}` +in `MvPolynomial ι R` is algebraically independent over `R`. -/ +theorem algebraicIndependent_polynomial_aeval_X + (f : ι → Polynomial R) (hf : ∀ i, Transcendental R (f i)) : + AlgebraicIndependent R fun i ↦ Polynomial.aeval (X i : MvPolynomial ι R) (f i) := by + set x := fun i ↦ Polynomial.aeval (X i : MvPolynomial ι R) (f i) + refine algebraicIndependent_of_finite_type' (C_injective _ _) fun t _ i hi ↦ ?_ + have hle : adjoin R (x '' t) ≤ supported R t := by + rw [Algebra.adjoin_le_iff, Set.image_subset_iff] + intro _ h + rw [Set.mem_preimage] + refine Algebra.adjoin_mono ?_ (Polynomial.aeval_mem_adjoin_singleton R _) + simp_rw [singleton_subset_iff, Set.mem_image_of_mem _ h] + exact (transcendental_supported_polynomial_aeval_X R hi (hf i)).of_tower_top_of_subalgebra_le hle + +end MvPolynomial + +/-- If `{x_i : A | i : ι}` is algebraically independent over `R`, and for each `i`, +`f_i : R[X]` is transcendental over `R`, then `{f_i(x_i) | i : ι}` is also +algebraically independent over `R`. -/ +theorem AlgebraicIndependent.polynomial_aeval_of_transcendental + (hx : AlgebraicIndependent R x) + {f : ι → Polynomial R} (hf : ∀ i, Transcendental R (f i)) : + AlgebraicIndependent R fun i ↦ Polynomial.aeval (x i) (f i) := by + convert aeval_of_algebraicIndependent hx (algebraicIndependent_polynomial_aeval_X _ hf) + rw [← AlgHom.comp_apply] + congr 1; ext1; simp From 3b0e6798409ef6c5a93238e93ebd6b432890c6a5 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Mon, 2 Dec 2024 12:55:53 +0000 Subject: [PATCH 467/829] chore(GroupTheory/Goursat): fix typos (#19686) Fix some minor typos and streamline a proof in the file on Goursat's lemma for groups. --- Mathlib/Algebra/Group/Graph.lean | 2 +- Mathlib/GroupTheory/Goursat.lean | 27 ++++++++++----------------- 2 files changed, 11 insertions(+), 18 deletions(-) diff --git a/Mathlib/Algebra/Group/Graph.lean b/Mathlib/Algebra/Group/Graph.lean index a3cf4bee29867..45c759e8ff68f 100644 --- a/Mathlib/Algebra/Group/Graph.lean +++ b/Mathlib/Algebra/Group/Graph.lean @@ -163,7 +163,7 @@ set_option linter.existingAttributeWarning false in attribute [to_additive existing] coe_graph graph_toSubmonoid @[to_additive] -lemma mem_graph {f : G →* H} {x : G × H} : x ∈ f.mgraph ↔ f x.1 = x.2 := .rfl +lemma mem_graph {f : G →* H} {x : G × H} : x ∈ f.graph ↔ f x.1 = x.2 := .rfl @[to_additive] lemma graph_eq_range_prod (f : G →* H) : f.graph = range ((id _).prod f) := by aesop diff --git a/Mathlib/GroupTheory/Goursat.lean b/Mathlib/GroupTheory/Goursat.lean index 90bc8ec652129..3864263abddfd 100644 --- a/Mathlib/GroupTheory/Goursat.lean +++ b/Mathlib/GroupTheory/Goursat.lean @@ -76,11 +76,13 @@ lemma mk_goursatFst_eq_iff_mk_goursatSnd_eq {x y : G × H} (hx : x ∈ I) (hy : · simpa [Prod.mul_def, Prod.div_def] using div_mem (mul_mem h hx) hy · simpa [Prod.mul_def, Prod.div_def] using div_mem (mul_mem h hy) hx +variable (I) in +@[to_additive AddSubgroup.goursatFst_prod_goursatSnd_le] lemma goursatFst_prod_goursatSnd_le : I.goursatFst.prod I.goursatSnd ≤ I := by rintro ⟨g, h⟩ ⟨hg, hh⟩ simpa using mul_mem (mem_goursatFst.1 hg) (mem_goursatSnd.1 hh) -/-- **Goursat's lemma** for a subgroup of with surjective projections. +/-- **Goursat's lemma** for a subgroup of a product with surjective projections. If `I` is a subgroup of `G × H` which projects fully on both factors, then there exist normal subgroups `M ≤ G` and `N ≤ H` such that `G' × H' ≤ I` and the image of `I` in `G ⧸ M × H ⧸ N` is the @@ -88,7 +90,7 @@ graph of an isomorphism `G ⧸ M ≃ H ⧸ N'`. `G'` and `H'` can be explicitly constructed as `I.goursatFst` and `I.goursatSnd` respectively. -/ @[to_additive -"**Goursat's lemma** for a subgroup of with surjective projections. +"**Goursat's lemma** for a subgroup of a product with surjective projections. If `I` is a subgroup of `G × H` which projects fully on both factors, then there exist normal subgroups `M ≤ G` and `N ≤ H` such that `G' × H' ≤ I` and the image of `I` in `G ⧸ M × H ⧸ N` is the @@ -111,8 +113,8 @@ lemma goursat_surjective : /-- **Goursat's lemma** for an arbitrary subgroup. If `I` is a subgroup of `G × H`, then there exist subgroups `G' ≤ G`, `H' ≤ H` and normal subgroups -`M ≤ G'` and `N ≤ H'` such that `M × N ≤ I` and the image of `I` in `G' ⧸ M × H' ⧸ N` is the graph -of an isomorphism `G ⧸ G' ≃ H ⧸ H'`. -/ +`M ⊴ G'` and `N ⊴ H'` such that `M × N ≤ I` and the image of `I` in `G' ⧸ M × H' ⧸ N` is the graph +of an isomorphism `G' ⧸ M ≃ H' ⧸ N`. -/ @[to_additive "**Goursat's lemma** for an arbitrary subgroup. @@ -161,18 +163,9 @@ lemma goursat : rintro h₁ hgh₁ g₁ hg₁h g₂ h₂ hg₂h₂ hP hQ simp only [Subtype.ext_iff] at hP hQ rwa [← hP, ← hQ] - · rintro ⟨⟨g, _⟩, ⟨h, _⟩⟩ hgh - simp only [MonoidHom.prodMap, MonoidHom.mem_ker, MonoidHom.prod_apply, MonoidHom.coe_comp, - QuotientGroup.coe_mk', MonoidHom.coe_fst, comp_apply, MonoidHom.coe_snd, Prod.mk_eq_one, - QuotientGroup.eq_one_iff, mem_goursatFst, MonoidHom.mem_range, Prod.mk.injEq, Subtype.exists, - Prod.exists, mem_goursatSnd] at hgh - rcases hgh with ⟨⟨g₁, h₁, hg₁h₁, hPQ₁⟩, ⟨g₂, h₂, hg₂h₂, hPQ₂⟩⟩ - simp only [Subtype.ext_iff] at hPQ₁ hPQ₂ - rcases hPQ₁ with ⟨rfl, rfl⟩ - rcases hPQ₂ with ⟨rfl, rfl⟩ - simp only [MonoidHom.mem_range, MonoidHom.prod_apply, Prod.mk.injEq, Subtype.exists, - Prod.exists] - refine ⟨g₁, h₂, ?_, rfl, rfl⟩ - simpa only [OneMemClass.coe_one, Prod.mk_mul_mk, mul_one, one_mul] using mul_mem hg₁h₁ hg₂h₂ + · convert goursatFst_prod_goursatSnd_le (P.prod Q).range + ext ⟨g, h⟩ + simp_rw [MonoidHom.mem_ker, MonoidHom.coe_prodMap, Prod.map_apply, Subgroup.mem_prod, + Prod.one_eq_mk, Prod.ext_iff, ← MonoidHom.mem_ker, QuotientGroup.ker_mk'] end Subgroup From e9404ad976f685168d45380b3510c7e2f319bbc8 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Mon, 2 Dec 2024 13:23:08 +0000 Subject: [PATCH 468/829] feat(NumberTheory/ModularForms/LevelOne) (#19479) add dimensions in non-positive weights. Co-authored-by: Chris Birkbeck Co-authored-by: David Loeffler --- Mathlib/NumberTheory/ModularForms/Basic.lean | 4 +- .../ModularForms/CongruenceSubgroups.lean | 3 + .../NumberTheory/ModularForms/LevelOne.lean | 88 ++++++++++++++++--- .../NumberTheory/ModularForms/QExpansion.lean | 12 ++- 4 files changed, 89 insertions(+), 18 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/Basic.lean b/Mathlib/NumberTheory/ModularForms/Basic.lean index 2a7ff22c6f898..67c2a36ff41f4 100644 --- a/Mathlib/NumberTheory/ModularForms/Basic.lean +++ b/Mathlib/NumberTheory/ModularForms/Basic.lean @@ -235,7 +235,6 @@ theorem mul_coe {k_1 k_2 : ℤ} {Γ : Subgroup SL(2, ℤ)} (f : ModularForm Γ k rfl /-- The constant function with value `x : ℂ` as a modular form of weight 0 and any level. -/ -@[simps! (config := .asFn) toFun toSlashInvariantForm] def const (x : ℂ) : ModularForm Γ 0 where toSlashInvariantForm := .const x holo' _ := mdifferentiableAt_const @@ -243,6 +242,9 @@ def const (x : ℂ) : ModularForm Γ 0 where simpa only [SlashInvariantForm.const_toFun, ModularForm.is_invariant_const] using atImInfty.const_boundedAtFilter x +@[simp] +lemma const_apply (x : ℂ) (τ : ℍ) : (const x : ModularForm Γ 0) τ = x := rfl + instance : One (ModularForm Γ 0) where one := { const 1 with toSlashInvariantForm := 1 } diff --git a/Mathlib/NumberTheory/ModularForms/CongruenceSubgroups.lean b/Mathlib/NumberTheory/ModularForms/CongruenceSubgroups.lean index e5b29a57f1dd4..ee91db811cdfa 100644 --- a/Mathlib/NumberTheory/ModularForms/CongruenceSubgroups.lean +++ b/Mathlib/NumberTheory/ModularForms/CongruenceSubgroups.lean @@ -64,6 +64,9 @@ theorem Gamma_one_top : Gamma 1 = ⊤ := by ext simp [eq_iff_true_of_subsingleton] +lemma mem_Gamma_one (γ : SL(2, ℤ)) : γ ∈ Γ(1) := by + simp only [Gamma_one_top, Subgroup.mem_top] + theorem Gamma_zero_bot : Gamma 0 = ⊥ := by ext simp only [Gamma_mem, coe_matrix_coe, Int.coe_castRingHom, map_apply, Int.cast_id, diff --git a/Mathlib/NumberTheory/ModularForms/LevelOne.lean b/Mathlib/NumberTheory/ModularForms/LevelOne.lean index 20ba662c6cc3b..5274e56478b89 100644 --- a/Mathlib/NumberTheory/ModularForms/LevelOne.lean +++ b/Mathlib/NumberTheory/ModularForms/LevelOne.lean @@ -3,8 +3,9 @@ Copyright (c) 2024 Chris Birkbeck. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Birkbeck -/ +import Mathlib.Analysis.Complex.AbsMax import Mathlib.NumberTheory.Modular -import Mathlib.NumberTheory.ModularForms.SlashInvariantForms +import Mathlib.NumberTheory.ModularForms.QExpansion /-! # Level one modular forms @@ -14,25 +15,84 @@ TODO: Add finite-dimensionality of these spaces of modular forms. -/ -open UpperHalfPlane ModularGroup SlashInvariantForm ModularForm Complex MatrixGroups +open UpperHalfPlane ModularGroup SlashInvariantForm ModularForm Complex + CongruenceSubgroup Real Function SlashInvariantFormClass ModularFormClass Periodic -lemma SlashInvariantForm.exists_one_half_le_im_and_norm_le {k : ℤ} (hk : k ≤ 0) {F : Type*} - [FunLike F ℍ ℂ] [SlashInvariantFormClass F ⊤ k] (f : F) (τ : ℍ) : - ∃ ξ : ℍ, 1/2 ≤ ξ.im ∧ ‖f τ‖ ≤ ‖f ξ‖ := +local notation "𝕢" => qParam + +variable {F : Type*} [FunLike F ℍ ℂ] {k : ℤ} + +namespace SlashInvariantForm + +variable [SlashInvariantFormClass F Γ(1) k] + +lemma exists_one_half_le_im_and_norm_le (hk : k ≤ 0) (f : F) (τ : ℍ) : + ∃ ξ : ℍ, 1 / 2 ≤ ξ.im ∧ ‖f τ‖ ≤ ‖f ξ‖ := let ⟨γ, hγ, hdenom⟩ := exists_one_half_le_im_smul_and_norm_denom_le τ - ⟨γ • τ, hγ, by simpa only [slash_action_eqn'' _ (show γ ∈ ⊤ by tauto), norm_mul, norm_zpow] - using le_mul_of_one_le_left (norm_nonneg _) <| + ⟨γ • τ, hγ, by simpa only [slash_action_eqn'' _ (mem_Gamma_one γ), + norm_mul, norm_zpow] using le_mul_of_one_le_left (norm_nonneg _) <| one_le_zpow_of_nonpos₀ (norm_pos_iff.2 (denom_ne_zero _ _)) hdenom hk⟩ +variable (k) in /-- If a constant function is modular of weight `k`, then either `k = 0`, or the constant is `0`. -/ -lemma SlashInvariantForm.wt_eq_zero_of_eq_const - {F : Type*} [FunLike F ℍ ℂ] (k : ℤ) [SlashInvariantFormClass F ⊤ k] - {f : F} {c : ℂ} (hf : ∀ τ, f τ = c) : k = 0 ∨ c = 0 := by - have hI := slash_action_eqn'' f (by tauto : ModularGroup.S ∈ ⊤) I - have h2I2 := slash_action_eqn'' f (by tauto : ModularGroup.S ∈ ⊤) ⟨2 * Complex.I, by simp⟩ - simp only [sl_moeb, hf, denom_S, coe_mk_subtype] at hI h2I2 +lemma wt_eq_zero_of_eq_const {f : F} {c : ℂ} (hf : ⇑f = Function.const _ c) : + k = 0 ∨ c = 0 := by + have hI := slash_action_eqn'' f (mem_Gamma_one S) I + have h2I2 := slash_action_eqn'' f (mem_Gamma_one S) ⟨2 * Complex.I, by norm_num⟩ + simp only [sl_moeb, hf, Function.const, denom_S, coe_mk_subtype] at hI h2I2 nth_rw 1 [h2I2] at hI simp only [mul_zpow, coe_I, mul_eq_mul_right_iff, mul_left_eq_self₀] at hI refine hI.imp_left (Or.casesOn · (fun H ↦ ?_) (False.elim ∘ zpow_ne_zero k I_ne_zero)) - rwa [← Complex.ofReal_ofNat, ← ofReal_zpow, ← ofReal_one, ofReal_inj, + rwa [← ofReal_ofNat, ← ofReal_zpow, ← ofReal_one, ofReal_inj, zpow_eq_one_iff_right₀ (by norm_num) (by norm_num)] at H + +end SlashInvariantForm + +namespace ModularFormClass + +variable [ModularFormClass F Γ(1) k] + +private theorem cuspFunction_eqOn_const_of_nonpos_wt (hk : k ≤ 0) (f : F) : + Set.EqOn (cuspFunction 1 f) (const ℂ (cuspFunction 1 f 0)) (Metric.ball 0 1) := by + refine eq_const_of_exists_le (fun q hq ↦ ?_) (exp_nonneg (-π)) ?_ (fun q hq ↦ ?_) + · exact (differentiableAt_cuspFunction 1 f (mem_ball_zero_iff.mp hq)).differentiableWithinAt + · simp only [exp_lt_one_iff, Left.neg_neg_iff, pi_pos] + · simp only [Metric.mem_closedBall, dist_zero_right] + rcases eq_or_ne q 0 with rfl | hq' + · refine ⟨0, by simpa only [norm_zero] using exp_nonneg _, le_rfl⟩ + · obtain ⟨ξ, hξ, hξ₂⟩ := exists_one_half_le_im_and_norm_le hk f + ⟨_, im_invQParam_pos_of_abs_lt_one Real.zero_lt_one (mem_ball_zero_iff.mp hq) hq'⟩ + exact ⟨_, abs_qParam_le_of_one_half_le_im hξ, + by simpa only [← eq_cuspFunction 1 f, Nat.cast_one, coe_mk_subtype, + qParam_right_inv one_ne_zero hq'] using hξ₂⟩ + +private theorem levelOne_nonpos_wt_const (hk : k ≤ 0) (f : F) : + ⇑f = Function.const _ (cuspFunction 1 f 0) := by + ext z + have hQ : 𝕢 1 z ∈ (Metric.ball 0 1) := by + simpa only [Metric.mem_ball, dist_zero_right, Complex.norm_eq_abs, neg_mul, mul_zero, div_one, + Real.exp_zero] using (abs_qParam_lt_iff zero_lt_one 0 z.1).mpr z.2 + simpa only [← eq_cuspFunction 1 f z, Nat.cast_one, Function.const_apply] using + (cuspFunction_eqOn_const_of_nonpos_wt hk f) hQ + +lemma levelOne_neg_weight_eq_zero (hk : k < 0) (f : F) : ⇑f = 0 := by + have hf := levelOne_nonpos_wt_const hk.le f + rcases wt_eq_zero_of_eq_const k hf with rfl | hf₀ + · exact (lt_irrefl _ hk).elim + · rw [hf, hf₀, const_zero] + +lemma levelOne_weight_zero_const [ModularFormClass F Γ(1) 0] (f : F) : + ∃ c, ⇑f = Function.const _ c := + ⟨_, levelOne_nonpos_wt_const le_rfl f⟩ + +end ModularFormClass + +lemma ModularForm.levelOne_weight_zero_rank_one : Module.rank ℂ (ModularForm Γ(1) 0) = 1 := by + refine rank_eq_one (const 1) (by simp [DFunLike.ne_iff]) fun g ↦ ?_ + obtain ⟨c', hc'⟩ := levelOne_weight_zero_const g + aesop + +lemma ModularForm.levelOne_neg_weight_rank_zero (hk : k < 0) : + Module.rank ℂ (ModularForm Γ(1) k) = 0 := by + refine rank_eq_zero_iff.mpr fun f ↦ ⟨_, one_ne_zero, ?_⟩ + simpa only [one_smul, ← DFunLike.coe_injective.eq_iff] using levelOne_neg_weight_eq_zero hk f diff --git a/Mathlib/NumberTheory/ModularForms/QExpansion.lean b/Mathlib/NumberTheory/ModularForms/QExpansion.lean index c236ef8443cf2..c8512f5b76f7e 100644 --- a/Mathlib/NumberTheory/ModularForms/QExpansion.lean +++ b/Mathlib/NumberTheory/ModularForms/QExpansion.lean @@ -20,9 +20,9 @@ application, we show that cusp forms decay exponentially to 0 as `im τ → ∞` * define the `q`-expansion as a formal power series -/ -open ModularForm Complex Filter Asymptotics UpperHalfPlane Function +open ModularForm Complex Filter UpperHalfPlane Function -open scoped Real Topology Manifold MatrixGroups CongruenceSubgroup +open scoped Real MatrixGroups CongruenceSubgroup noncomputable section @@ -36,7 +36,13 @@ theorem Function.Periodic.im_invQParam_pos_of_abs_lt_one 0 < im (Periodic.invQParam h q) := im_invQParam .. ▸ mul_pos_of_neg_of_neg (div_neg_of_neg_of_pos (neg_lt_zero.mpr hh) Real.two_pi_pos) - ((Real.log_neg_iff (Complex.abs.pos hq_ne)).mpr hq) + ((Real.log_neg_iff (abs.pos hq_ne)).mpr hq) + +lemma Function.Periodic.abs_qParam_le_of_one_half_le_im {ξ : ℂ} (hξ : 1 / 2 ≤ ξ.im) : + ‖𝕢 1 ξ‖ ≤ rexp (-π) := by + rwa [Periodic.qParam, ofReal_one, div_one, norm_eq_abs, abs_exp, Real.exp_le_exp, + mul_right_comm, mul_I_re, neg_le_neg_iff, ← ofReal_ofNat, ← ofReal_mul, im_ofReal_mul, + mul_comm _ π, mul_assoc, le_mul_iff_one_le_right Real.pi_pos, ← div_le_iff₀' two_pos] namespace SlashInvariantFormClass From 62a0adf40aa51bc5fc8c010bc99a7fa8031b1e8a Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Mon, 2 Dec 2024 13:35:34 +0000 Subject: [PATCH 469/829] feat: roots have non-zero length wrt the canonical form of a finite crystallographic root pairing (#19645) We also relocate some other unrelated lemmas in order to relax the assumptions on their scalars, and add fill out some loosely-related API. The only mathematical result is `RootPairing.rootForm_apply_root_self_ne_zero`. --- Mathlib/LinearAlgebra/PerfectPairing.lean | 17 +++++ Mathlib/LinearAlgebra/RootSystem/Basic.lean | 3 +- Mathlib/LinearAlgebra/RootSystem/Defs.lean | 6 ++ .../RootSystem/Finite/CanonicalBilinear.lean | 75 ++++++++++++++++--- .../RootSystem/Finite/Nondegenerate.lean | 6 +- 5 files changed, 89 insertions(+), 18 deletions(-) diff --git a/Mathlib/LinearAlgebra/PerfectPairing.lean b/Mathlib/LinearAlgebra/PerfectPairing.lean index ca394d510fa52..46a72d4138489 100644 --- a/Mathlib/LinearAlgebra/PerfectPairing.lean +++ b/Mathlib/LinearAlgebra/PerfectPairing.lean @@ -165,6 +165,23 @@ include p in theorem reflexive_right : IsReflexive R N := p.flip.reflexive_left +instance : EquivLike (PerfectPairing R M N) M (Dual R N) where + coe p := p.toDualLeft + inv p := p.toDualLeft.symm + left_inv p x := LinearEquiv.symm_apply_apply _ _ + right_inv p x := LinearEquiv.apply_symm_apply _ _ + coe_injective' p q h h' := by + cases p + cases q + simp only [mk.injEq] + ext m n + simp only [DFunLike.coe_fn_eq] at h + exact LinearMap.congr_fun (LinearEquiv.congr_fun h m) n + +instance : LinearEquivClass (PerfectPairing R M N) R M (Dual R N) where + map_add p m₁ m₂ := p.toLin.map_add m₁ m₂ + map_smulₛₗ p t m := p.toLin.map_smul t m + include p in theorem finrank_eq [Module.Finite R M] [Module.Free R M] : finrank R M = finrank R N := diff --git a/Mathlib/LinearAlgebra/RootSystem/Basic.lean b/Mathlib/LinearAlgebra/RootSystem/Basic.lean index 4aba11c944d5e..5a32efb2dc700 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Basic.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Basic.lean @@ -161,7 +161,8 @@ private lemma coroot_eq_coreflection_of_root_eq' [CharZero R] [NoZeroSMulDivisor rw [← hl] have hkl : (p.flip (coroot l)) (root k) = 2 := by simp only [hl, preReflection_apply, hk, PerfectPairing.flip_apply_apply, map_sub, hp j, - map_smul, smul_eq_mul, hp i, mul_sub, sα, α, α', β, mul_two, mul_add] + map_smul, smul_eq_mul, hp i, mul_sub, sα, α, α', β, mul_two, mul_add, LinearMap.sub_apply, + LinearMap.smul_apply] rw [mul_comm (p (root i) (coroot j))] abel suffices p.flip (coroot k) = p.flip (coroot l) from p.bijectiveRight.injective this diff --git a/Mathlib/LinearAlgebra/RootSystem/Defs.lean b/Mathlib/LinearAlgebra/RootSystem/Defs.lean index d9ab0eb932455..2f766ce1ab82f 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Defs.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Defs.lean @@ -359,6 +359,12 @@ lemma isCrystallographic_iff : · simpa [AddSubgroup.mem_zmultiples_iff] using h i (mem_range_self j) · simpa [← hj, AddSubgroup.mem_zmultiples_iff] using h i j +variable {P} in +lemma IsCrystallographic.flip (h : P.IsCrystallographic) : + P.flip.IsCrystallographic := by + rw [isCrystallographic_iff, forall_comm] + exact P.isCrystallographic_iff.mp h + /-- A root pairing is said to be reduced if any linearly dependent pair of roots is related by a sign. -/ def IsReduced : Prop := diff --git a/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean b/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean index 1f7676b71d4ad..9917885d01294 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean @@ -55,7 +55,11 @@ namespace RootPairing section CommRing variable [Fintype ι] [CommRing R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] -(P : RootPairing ι R M N) + (P : RootPairing ι R M N) + +instance : Module.Finite R P.rootSpan := Finite.span_of_finite R <| finite_range P.root + +instance : Module.Finite R P.corootSpan := Finite.span_of_finite R <| finite_range P.coroot /-- An invariant linear map from weight space to coweight space. -/ def Polarization : M →ₗ[R] N := @@ -87,9 +91,33 @@ def CorootForm : LinearMap.BilinForm R N := ∑ i, (P.root' i).smulRight (P.root' i) lemma rootForm_apply_apply (x y : M) : P.RootForm x y = - ∑ (i : ι), P.coroot' i x * P.coroot' i y := by + ∑ i, P.coroot' i x * P.coroot' i y := by simp [RootForm] +lemma corootForm_apply_apply (x y : N) : P.CorootForm x y = + ∑ i, P.root' i x * P.root' i y := by + simp [CorootForm] + +lemma flip_comp_polarization_eq_rootForm : + P.flip.toLin ∘ₗ P.Polarization = P.RootForm := by + ext; simp [rootForm_apply_apply, RootPairing.flip] + +lemma self_comp_coPolarization_eq_corootForm : + P.toLin ∘ₗ P.CoPolarization = P.CorootForm := by + ext; simp [corootForm_apply_apply] + +lemma polarization_apply_eq_zero_iff (m : M) : + P.Polarization m = 0 ↔ P.RootForm m = 0 := by + rw [← flip_comp_polarization_eq_rootForm] + refine ⟨fun h ↦ by simp [h], fun h ↦ ?_⟩ + change P.toDualRight (P.Polarization m) = 0 at h + simp only [EmbeddingLike.map_eq_zero_iff] at h + exact h + +lemma coPolarization_apply_eq_zero_iff (n : N) : + P.CoPolarization n = 0 ↔ P.CorootForm n = 0 := + P.flip.polarization_apply_eq_zero_iff n + lemma rootForm_symmetric : LinearMap.IsSymm P.RootForm := by simp [LinearMap.IsSymm, mul_comm, rootForm_apply_apply] @@ -143,6 +171,39 @@ theorem range_polarization_domRestrict_le_span_coroot : use fun i => (P.toPerfectPairing x) (P.coroot i) simp +lemma prod_rootForm_smul_coroot_mem_range_domRestrict (i : ι) : + (∏ a : ι, P.RootForm (P.root a) (P.root a)) • P.coroot i ∈ + LinearMap.range (P.Polarization.domRestrict (P.rootSpan)) := by + obtain ⟨c, hc⟩ := Finset.dvd_prod_of_mem (fun a ↦ P.RootForm (P.root a) (P.root a)) + (Finset.mem_univ i) + rw [hc, mul_comm, mul_smul, rootForm_self_smul_coroot] + refine LinearMap.mem_range.mpr ?_ + use ⟨(c • 2 • P.root i), by aesop⟩ + simp + +section IsCrystallographic + +variable [CharZero R] (h : P.IsCrystallographic) (i : ι) +include h + +lemma rootForm_apply_root_self_ne_zero : + P.RootForm (P.root i) (P.root i) ≠ 0 := by + choose z hz using P.isCrystallographic_iff.mp h i + simp only [rootForm_apply_apply, PerfectPairing.flip_apply_apply, root_coroot_eq_pairing, ← hz] + suffices 0 < ∑ i, z i * z i by norm_cast; exact this.ne' + refine Finset.sum_pos' (fun i _ ↦ mul_self_nonneg (z i)) ⟨i, Finset.mem_univ i, ?_⟩ + have hzi : z i = 2 := by + specialize hz i + rw [pairing_same] at hz + norm_cast at hz + simp [hzi] + +lemma corootForm_apply_coroot_self_ne_zero : + P.CorootForm (P.coroot i) (P.coroot i) ≠ 0 := + P.flip.rootForm_apply_root_self_ne_zero h.flip i + +end IsCrystallographic + end CommRing section LinearOrderedCommRing @@ -172,16 +233,6 @@ lemma prod_rootForm_root_self_pos : 0 < ∏ i, P.RootForm (P.root i) (P.root i) := Finset.prod_pos fun i _ => rootForm_root_self_pos P i -lemma prod_rootForm_smul_coroot_mem_range_domRestrict (i : ι) : - (∏ a : ι, P.RootForm (P.root a) (P.root a)) • P.coroot i ∈ - LinearMap.range (P.Polarization.domRestrict (P.rootSpan)) := by - obtain ⟨c, hc⟩ := Finset.dvd_prod_of_mem (fun a ↦ P.RootForm (P.root a) (P.root a)) - (Finset.mem_univ i) - rw [hc, mul_comm, mul_smul, rootForm_self_smul_coroot] - refine LinearMap.mem_range.mpr ?_ - use ⟨(c • 2 • P.root i), by aesop⟩ - simp - end LinearOrderedCommRing end RootPairing diff --git a/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean b/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean index 1efe86c29f24f..821094cd58377 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean @@ -53,17 +53,13 @@ namespace RootPairing variable {ι R M N : Type*} variable [Fintype ι] [LinearOrderedCommRing R] [AddCommGroup M] [Module R M] [AddCommGroup N] -[Module R N] (P : RootPairing ι R M N) + [Module R N] (P : RootPairing ι R M N) lemma rootForm_rootPositive : IsRootPositive P P.RootForm where zero_lt_apply_root i := P.rootForm_root_self_pos i symm := P.rootForm_symmetric apply_reflection_eq := P.rootForm_reflection_reflection_apply -instance : Module.Finite R P.rootSpan := Finite.span_of_finite R <| finite_range P.root - -instance : Module.Finite R P.corootSpan := Finite.span_of_finite R <| finite_range P.coroot - @[simp] lemma finrank_rootSpan_map_polarization_eq_finrank_corootSpan : finrank R (P.rootSpan.map P.Polarization) = finrank R P.corootSpan := by From 4587215fb214d15564077921d596506998192814 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 2 Dec 2024 13:47:31 +0000 Subject: [PATCH 470/829] =?UTF-8?q?feat(Finset):=20`(s=20=E2=88=AA=20t).No?= =?UTF-8?q?nempty=20=E2=86=94=20s.Nonempty=20=E2=88=A8=20t.Nonempty`=20(#1?= =?UTF-8?q?9609)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From GrowthInGroups (LeanCamCombi) --- Mathlib/Data/Finset/Lattice/Lemmas.lean | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/Mathlib/Data/Finset/Lattice/Lemmas.lean b/Mathlib/Data/Finset/Lattice/Lemmas.lean index ad6b5a9afb949..c5783884b4b5a 100644 --- a/Mathlib/Data/Finset/Lattice/Lemmas.lean +++ b/Mathlib/Data/Finset/Lattice/Lemmas.lean @@ -155,7 +155,9 @@ lemma inter_singleton {a : α} {s : Finset α} : s ∩ {a} = if a ∈ s then {a} else ∅ := by split_ifs with h <;> simp [h] -lemma union_eq_empty : s ∪ t = ∅ ↔ s = ∅ ∧ t = ∅ := sup_eq_bot_iff +@[simp] lemma union_eq_empty : s ∪ t = ∅ ↔ s = ∅ ∧ t = ∅ := sup_eq_bot_iff +@[simp] lemma union_nonempty : (s ∪ t).Nonempty ↔ s.Nonempty ∨ t.Nonempty := + mod_cast Set.union_nonempty (α := α) (s := s) (t := t) theorem insert_union_comm (s t : Finset α) (a : α) : insert a s ∪ t = s ∪ insert a t := by rw [insert_union, union_insert] From 7cf807751deab8d4943d867898dc1b31d61b746a Mon Sep 17 00:00:00 2001 From: "Filippo A. E. Nuccio" Date: Mon, 2 Dec 2024 14:24:32 +0000 Subject: [PATCH 471/829] docs(RootSystem/Defs): add more details to the `doc` for root pairings. (#19683) Add a comparison with the "classical" literature in the definition of a root pairing. --- Mathlib/LinearAlgebra/RootSystem/Defs.lean | 6 +++++- 1 file changed, 5 insertions(+), 1 deletion(-) diff --git a/Mathlib/LinearAlgebra/RootSystem/Defs.lean b/Mathlib/LinearAlgebra/RootSystem/Defs.lean index 2f766ce1ab82f..3970a7e11cf75 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Defs.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Defs.lean @@ -92,7 +92,11 @@ structure RootPairing extends PerfectPairing R M N where /-- A parametrized family of dual vectors, called coroots. -/ coroot : ι ↪ N root_coroot_two : ∀ i, toLin (root i) (coroot i) = 2 - /-- A parametrized family of permutations, induced by reflection. -/ + /-- A parametrized family of permutations, induced by reflections. This corresponds to the + classical requirement that the symmetry attached to each root (later defined in + `RootPairing.reflection`) leave the whole set of roots stable: as explained above, we + formalize this stability by fixing the image of the roots through each reflection (whence the + permutation); and similarly for coroots. -/ reflection_perm : ι → (ι ≃ ι) reflection_perm_root : ∀ i j, root j - toPerfectPairing (root j) (coroot i) • root i = root (reflection_perm i j) From 1e85f24d1fbaf6b4990cd795d106377982a19971 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 2 Dec 2024 14:42:28 +0000 Subject: [PATCH 472/829] feat: `gcongr` lemmas for `Additive`/`Multiplicative` (#18959) From FLT Co-authored-by: Javier Lopez-Contreras --- Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean | 10 ++++++++++ Mathlib/MeasureTheory/OuterMeasure/Basic.lean | 2 +- 2 files changed, 11 insertions(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean index 761ae23ae5a20..483877e4fee6e 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean @@ -86,6 +86,11 @@ theorem toMul_le {a b : Additive α} : a.toMul ≤ b.toMul ↔ a ≤ b := theorem toMul_lt {a b : Additive α} : a.toMul < b.toMul ↔ a < b := Iff.rfl +@[gcongr] alias ⟨_, toMul_mono⟩ := toMul_le +@[gcongr] alias ⟨_, ofMul_mono⟩ := ofMul_le +@[gcongr] alias ⟨_, toMul_strictMono⟩ := toMul_lt +@[gcongr] alias ⟨_, foMul_strictMono⟩ := ofMul_lt + end Additive namespace Multiplicative @@ -108,4 +113,9 @@ theorem toAdd_le {a b : Multiplicative α} : a.toAdd ≤ b.toAdd ↔ a ≤ b := theorem toAdd_lt {a b : Multiplicative α} : a.toAdd < b.toAdd ↔ a < b := Iff.rfl +@[gcongr] alias ⟨_, toAdd_mono⟩ := toAdd_le +@[gcongr] alias ⟨_, ofAdd_mono⟩ := ofAdd_le +@[gcongr] alias ⟨_, toAdd_strictMono⟩ := toAdd_lt +@[gcongr] alias ⟨_, ofAdd_strictMono⟩ := ofAdd_lt + end Multiplicative diff --git a/Mathlib/MeasureTheory/OuterMeasure/Basic.lean b/Mathlib/MeasureTheory/OuterMeasure/Basic.lean index 42207c8540857..24180bd2709cb 100644 --- a/Mathlib/MeasureTheory/OuterMeasure/Basic.lean +++ b/Mathlib/MeasureTheory/OuterMeasure/Basic.lean @@ -61,7 +61,7 @@ theorem measure_iUnion_le [Countable ι] (s : ι → Set α) : μ (⋃ i, s i) μ (⋃ i, t i) = μ (⋃ i, disjointed t i) := by rw [iUnion_disjointed] _ ≤ ∑' i, μ (disjointed t i) := OuterMeasureClass.measure_iUnion_nat_le _ _ (disjoint_disjointed _) - _ ≤ ∑' i, μ (t i) := by gcongr; apply disjointed_subset + _ ≤ ∑' i, μ (t i) := by gcongr; exact disjointed_subset .. theorem measure_biUnion_le {I : Set ι} (μ : F) (hI : I.Countable) (s : ι → Set α) : μ (⋃ i ∈ I, s i) ≤ ∑' i : I, μ (s i) := by From 0429dca7fb0f2703f1c4f86348c31afc5bd4c55c Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 2 Dec 2024 15:15:10 +0000 Subject: [PATCH 473/829] chore: defer a lemma, to avoid needing Cardinal in LinearIndependent (#19626) --- .../Algebra/Category/ModuleCat/Adjunctions.lean | 2 ++ Mathlib/Analysis/Convex/Cone/Basic.lean | 1 + Mathlib/Analysis/LocallyConvex/Polar.lean | 1 - Mathlib/LinearAlgebra/Dimension/Basic.lean | 1 + Mathlib/LinearAlgebra/Dimension/Finite.lean | 13 +++++++++++++ Mathlib/LinearAlgebra/Finsupp/VectorSpace.lean | 1 - Mathlib/LinearAlgebra/FreeModule/Basic.lean | 9 +++++---- Mathlib/LinearAlgebra/LinearIndependent.lean | 16 +--------------- Mathlib/RingTheory/TensorProduct/Free.lean | 2 ++ 9 files changed, 25 insertions(+), 21 deletions(-) diff --git a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean index e23eebad47930..12126e451178b 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean @@ -15,6 +15,8 @@ the forgetful functor from `R`-modules to types. -/ +assert_not_exists Cardinal + noncomputable section open CategoryTheory diff --git a/Mathlib/Analysis/Convex/Cone/Basic.lean b/Mathlib/Analysis/Convex/Cone/Basic.lean index e4fb40b9a90ef..303a1a919f4f2 100644 --- a/Mathlib/Analysis/Convex/Cone/Basic.lean +++ b/Mathlib/Analysis/Convex/Cone/Basic.lean @@ -39,6 +39,7 @@ While `Convex 𝕜` is a predicate on sets, `ConvexCone 𝕜 E` is a bundled con assert_not_exists NormedSpace assert_not_exists Real +assert_not_exists Cardinal open Set LinearMap diff --git a/Mathlib/Analysis/LocallyConvex/Polar.lean b/Mathlib/Analysis/LocallyConvex/Polar.lean index dba80d95bd857..af5dc28df843c 100644 --- a/Mathlib/Analysis/LocallyConvex/Polar.lean +++ b/Mathlib/Analysis/LocallyConvex/Polar.lean @@ -34,7 +34,6 @@ any bilinear form `B : E →ₗ[𝕜] F →ₗ[𝕜] 𝕜`, where `𝕜` is a no polar -/ - variable {𝕜 E F : Type*} open Topology diff --git a/Mathlib/LinearAlgebra/Dimension/Basic.lean b/Mathlib/LinearAlgebra/Dimension/Basic.lean index 16bbd326c3112..5e63e9aa3eb1a 100644 --- a/Mathlib/LinearAlgebra/Dimension/Basic.lean +++ b/Mathlib/LinearAlgebra/Dimension/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Johannes Hölzl, Sander Dahmen, Kim Morrison -/ import Mathlib.LinearAlgebra.LinearIndependent +import Mathlib.SetTheory.Cardinal.Basic /-! # Dimension of modules and vector spaces diff --git a/Mathlib/LinearAlgebra/Dimension/Finite.lean b/Mathlib/LinearAlgebra/Dimension/Finite.lean index 6c000a000aaf9..3ae92933cf238 100644 --- a/Mathlib/LinearAlgebra/Dimension/Finite.lean +++ b/Mathlib/LinearAlgebra/Dimension/Finite.lean @@ -29,6 +29,19 @@ attribute [local instance] nontrivial_of_invariantBasisNumber open Basis Cardinal Function Module Set Submodule +/-- If every finite set of linearly independent vectors has cardinality at most `n`, +then the same is true for arbitrary sets of linearly independent vectors. +-/ +theorem linearIndependent_bounded_of_finset_linearIndependent_bounded {n : ℕ} + (H : ∀ s : Finset M, (LinearIndependent R fun i : s => (i : M)) → s.card ≤ n) : + ∀ s : Set M, LinearIndependent R ((↑) : s → M) → #s ≤ n := by + intro s li + apply Cardinal.card_le_of + intro t + rw [← Finset.card_map (Embedding.subtype s)] + apply H + apply linearIndependent_finset_map_embedding_subtype _ li + theorem rank_le {n : ℕ} (H : ∀ s : Finset M, (LinearIndependent R fun i : s => (i : M)) → s.card ≤ n) : Module.rank R M ≤ n := by diff --git a/Mathlib/LinearAlgebra/Finsupp/VectorSpace.lean b/Mathlib/LinearAlgebra/Finsupp/VectorSpace.lean index 55e82e68d64d0..8fd119ba66b51 100644 --- a/Mathlib/LinearAlgebra/Finsupp/VectorSpace.lean +++ b/Mathlib/LinearAlgebra/Finsupp/VectorSpace.lean @@ -19,7 +19,6 @@ This file contains results on the `R`-module structure on functions of finite su noncomputable section open Set LinearMap Submodule -open scoped Cardinal universe u v w diff --git a/Mathlib/LinearAlgebra/FreeModule/Basic.lean b/Mathlib/LinearAlgebra/FreeModule/Basic.lean index 99e69913af94b..78dc32b0099f9 100644 --- a/Mathlib/LinearAlgebra/FreeModule/Basic.lean +++ b/Mathlib/LinearAlgebra/FreeModule/Basic.lean @@ -7,6 +7,7 @@ import Mathlib.Data.Finsupp.Fintype import Mathlib.Data.Matrix.Defs import Mathlib.LinearAlgebra.Basis.Basic import Mathlib.LinearAlgebra.TensorProduct.Basis +import Mathlib.Logic.Small.Basic /-! # Free modules @@ -37,6 +38,10 @@ variable [Semiring R] [AddCommMonoid M] [Module R M] class Module.Free : Prop where exists_basis : Nonempty <| (I : Type v) × Basis I R M +theorem Module.free_iff_set : Module.Free R M ↔ ∃ S : Set M, Nonempty (Basis S R M) := + ⟨fun h => ⟨Set.range h.exists_basis.some.2, ⟨Basis.reindexRange h.exists_basis.some.2⟩⟩, + fun ⟨S, hS⟩ => ⟨nonempty_sigma.2 ⟨S, hS⟩⟩⟩ + /-- If `M` fits in universe `w`, then freeness is equivalent to existence of a basis in that universe. @@ -49,10 +54,6 @@ theorem Module.free_def [Small.{w,v} M] : ⟨(Basis.reindexRange h.exists_basis.some.2).reindex (equivShrink _)⟩⟩, fun h => ⟨(nonempty_sigma.2 h).map fun ⟨_, b⟩ => ⟨Set.range b, b.reindexRange⟩⟩⟩ -theorem Module.free_iff_set : Module.Free R M ↔ ∃ S : Set M, Nonempty (Basis S R M) := - ⟨fun h => ⟨Set.range h.exists_basis.some.2, ⟨Basis.reindexRange h.exists_basis.some.2⟩⟩, - fun ⟨S, hS⟩ => ⟨nonempty_sigma.2 ⟨S, hS⟩⟩⟩ - variable {R M} theorem Module.Free.of_basis {ι : Type w} (b : Basis ι R M) : Module.Free R M := diff --git a/Mathlib/LinearAlgebra/LinearIndependent.lean b/Mathlib/LinearAlgebra/LinearIndependent.lean index eedfde103f8a0..88c16433bb3f7 100644 --- a/Mathlib/LinearAlgebra/LinearIndependent.lean +++ b/Mathlib/LinearAlgebra/LinearIndependent.lean @@ -7,7 +7,6 @@ import Mathlib.Algebra.BigOperators.Fin import Mathlib.Data.Set.Subsingleton import Mathlib.Lean.Expr.ExtraRecognizers import Mathlib.LinearAlgebra.Prod -import Mathlib.SetTheory.Cardinal.Basic import Mathlib.Tactic.FinCases import Mathlib.Tactic.LinearCombination import Mathlib.Tactic.Module @@ -82,13 +81,12 @@ linearly dependent, linear dependence, linearly independent, linear independence -/ +assert_not_exists Cardinal noncomputable section open Function Set Submodule -open Cardinal - universe u' u variable {ι : Type u'} {ι' : Type*} {R : Type*} {K : Type*} @@ -402,18 +400,6 @@ theorem linearIndependent_finset_map_embedding_subtype (s : Set M) obtain ⟨b, _hb, rfl⟩ := hy simp only [f, imp_self, Subtype.mk_eq_mk] -/-- If every finite set of linearly independent vectors has cardinality at most `n`, -then the same is true for arbitrary sets of linearly independent vectors. --/ -theorem linearIndependent_bounded_of_finset_linearIndependent_bounded {n : ℕ} - (H : ∀ s : Finset M, (LinearIndependent R fun i : s => (i : M)) → s.card ≤ n) : - ∀ s : Set M, LinearIndependent R ((↑) : s → M) → #s ≤ n := by - intro s li - apply Cardinal.card_le_of - intro t - rw [← Finset.card_map (Embedding.subtype s)] - apply H - apply linearIndependent_finset_map_embedding_subtype _ li section Subtype diff --git a/Mathlib/RingTheory/TensorProduct/Free.lean b/Mathlib/RingTheory/TensorProduct/Free.lean index a857cfdf874c6..a58fc59512f2a 100644 --- a/Mathlib/RingTheory/TensorProduct/Free.lean +++ b/Mathlib/RingTheory/TensorProduct/Free.lean @@ -22,6 +22,8 @@ and deducde that `Module.Free` is stable under base change. -/ +assert_not_exists Cardinal + suppress_compilation open scoped TensorProduct From 612f6c3a580b19d6e5dcb02ef1d246e65f4f4c73 Mon Sep 17 00:00:00 2001 From: Daniel Weber Date: Mon, 2 Dec 2024 15:15:12 +0000 Subject: [PATCH 474/829] feat(Algebra/BigOperators/Ring): add `Multiset.sum_map_div` lemma (#19630) --- Mathlib/Algebra/BigOperators/Ring.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/Mathlib/Algebra/BigOperators/Ring.lean b/Mathlib/Algebra/BigOperators/Ring.lean index b7468a17eabc1..6e86cf047cd5e 100644 --- a/Mathlib/Algebra/BigOperators/Ring.lean +++ b/Mathlib/Algebra/BigOperators/Ring.lean @@ -274,6 +274,10 @@ end CommRing section DivisionSemiring variable [DivisionSemiring α] +lemma _root_.Multiset.sum_map_div {s : Multiset ι} {f : ι → α} {a : α} : + (s.map (fun x ↦ f x / a)).sum = (s.map f).sum / a := by + simp only [div_eq_mul_inv, Multiset.sum_map_mul_right] + lemma sum_div (s : Finset ι) (f : ι → α) (a : α) : (∑ i ∈ s, f i) / a = ∑ i ∈ s, f i / a := by simp only [div_eq_mul_inv, sum_mul] From 5255add1b2ea28533d99d582c7cfcc20ad22939a Mon Sep 17 00:00:00 2001 From: grunweg Date: Mon, 2 Dec 2024 15:15:13 +0000 Subject: [PATCH 475/829] feat(check-yaml,yaml-check): various small tweaks (#19688) Extracted from #19634. --- scripts/check-yaml.lean | 23 ++++++++++------------- scripts/yaml_check.py | 12 +++++++++--- 2 files changed, 19 insertions(+), 16 deletions(-) diff --git a/scripts/check-yaml.lean b/scripts/check-yaml.lean index 89147aeed4c13..731c46a1cabab 100644 --- a/scripts/check-yaml.lean +++ b/scripts/check-yaml.lean @@ -11,7 +11,7 @@ import Mathlib.Tactic.ToExpr This assumes `yaml_check.py` has first translated these to `json` files. -It verifies that the referenced declarations exist. +It verifies that the referenced declarations exist, and prints an error otherwise. -/ open IO.FS Lean Lean.Elab @@ -23,18 +23,15 @@ def readJsonFile (α) [FromJson α] (path : System.FilePath) : IO α := do let _ : MonadExceptOf String IO := ⟨throw ∘ IO.userError, fun x _ => x⟩ liftExcept <| fromJson? <|← liftExcept <| Json.parse <|← IO.FS.readFile path -def databases : List (String × String) := - ["undergrad", "overview", "100"].map fun dir => - (dir ++ ".json", - s!"Entries in `docs/{dir}.yaml` refer to declarations that don't exist. \ - Please correct the following:") +def databases : List String := ["undergrad", "overview", "100"] -def processDb (decls : ConstMap) : String × String → IO Bool -| (file, msg) => do - let lines := ← readJsonFile DBFile file - let missing := lines.filter (fun l => !(decls.contains l.2)) +def processDb (decls : ConstMap) : String → IO Bool +| file => do + let lines ← readJsonFile DBFile s!"{file}.json" + let missing := lines.filter (fun l ↦ !(decls.contains l.2)) if 0 < missing.size then - IO.println msg + IO.println s!"Entries in `docs/{file}.yaml` refer to {missing.size} declaration(s) that don't exist. \ + Please correct the following:" for p in missing do IO.println s!" {p.1}: {p.2}" IO.println "" @@ -45,7 +42,7 @@ def processDb (decls : ConstMap) : String × String → IO Bool unsafe def main : IO Unit := do CoreM.withImportModules #[`Mathlib, `Archive] (searchPath := compile_time_search_path%) (trustLevel := 1024) do - let decls := (←getEnv).constants - let results ← databases.mapM (fun p => processDb decls p) + let decls := (← getEnv).constants + let results ← databases.mapM (fun p ↦ processDb decls p) if results.any id then IO.Process.exit 1 diff --git a/scripts/yaml_check.py b/scripts/yaml_check.py index 5d6ee74ddfa4f..4e05448995288 100644 --- a/scripts/yaml_check.py +++ b/scripts/yaml_check.py @@ -1,7 +1,13 @@ """ -This file is copied from the mathlib3 file of the same name. -It reads in the three yaml files, and translates them to simpler json files that are easier to -process in Lean. +This file reads in the three yaml files, and translates them to simpler json files +that are easier to process in Lean. + +Usage: + python3 yaml_check.py <100.yaml> + +(Each is the path to a yaml file containing information about the respective class + of theorems. The order of these files is important.) + """ from typing import Dict, Optional, Union, Tuple, List import yaml From 09ceeef3b669670fa610a5da84abbfa7e8c92dbd Mon Sep 17 00:00:00 2001 From: "Tristan F.-R." Date: Mon, 2 Dec 2024 15:59:09 +0000 Subject: [PATCH 476/829] feat(Logic/IsEmpty): Left/Right/BiTotal truthiness on emptiness (#19555) Adds left, right, and bi_total_empty, and golfs `identical_of_isEmpty` in `PGame` as a result. Co-authored-by: Tristan F. --- Mathlib/Logic/IsEmpty.lean | 17 +++++++++++++++++ Mathlib/SetTheory/Game/PGame.lean | 2 +- 2 files changed, 18 insertions(+), 1 deletion(-) diff --git a/Mathlib/Logic/IsEmpty.lean b/Mathlib/Logic/IsEmpty.lean index b46a933890f1c..aad5c6c29d0ee 100644 --- a/Mathlib/Logic/IsEmpty.lean +++ b/Mathlib/Logic/IsEmpty.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ import Mathlib.Logic.Function.Basic +import Mathlib.Logic.Relator /-! # Types that are empty @@ -204,3 +205,19 @@ variable {α} theorem Function.extend_of_isEmpty [IsEmpty α] (f : α → β) (g : α → γ) (h : β → γ) : Function.extend f g h = h := funext fun _ ↦ (Function.extend_apply' _ _ _) fun ⟨a, _⟩ ↦ isEmptyElim a + +open Relator + +variable {α β : Type*} (R : α → β → Prop) + +@[simp] +theorem leftTotal_empty [IsEmpty α] : LeftTotal R := by + simp only [LeftTotal, IsEmpty.forall_iff] + +@[simp] +theorem rightTotal_empty [IsEmpty β] : RightTotal R := by + simp only [RightTotal, IsEmpty.forall_iff] + +@[simp] +theorem biTotal_empty [IsEmpty α] [IsEmpty β] : BiTotal R := + ⟨leftTotal_empty R, rightTotal_empty R⟩ diff --git a/Mathlib/SetTheory/Game/PGame.lean b/Mathlib/SetTheory/Game/PGame.lean index 4b18cfd466466..d1e221f7c4878 100644 --- a/Mathlib/SetTheory/Game/PGame.lean +++ b/Mathlib/SetTheory/Game/PGame.lean @@ -392,7 +392,7 @@ theorem moveRight_memᵣ (x : PGame) (b) : x.moveRight b ∈ᵣ x := ⟨_, .rfl theorem identical_of_isEmpty (x y : PGame) [IsEmpty x.LeftMoves] [IsEmpty x.RightMoves] [IsEmpty y.LeftMoves] [IsEmpty y.RightMoves] : x ≡ y := - identical_iff.2 <| by simp [Relator.BiTotal, Relator.LeftTotal, Relator.RightTotal] + identical_iff.2 (by simp [biTotal_empty]) /-- `Identical` as a `Setoid`. -/ def identicalSetoid : Setoid PGame := From ceca16705b88a393e6e920556c405975ed8e3b48 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 2 Dec 2024 16:40:47 +0000 Subject: [PATCH 477/829] =?UTF-8?q?feat(SetTheory/Ordinal/Basic):=20notati?= =?UTF-8?q?on=20`typeLT=20=CE=B1=20=3D=20type=20(=CE=B1=20:=3D=20=CE=B1)?= =?UTF-8?q?=20(=C2=B7=20<=20=C2=B7)`=20(#18235)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/SetTheory/Cardinal/Aleph.lean | 2 +- Mathlib/SetTheory/Ordinal/Basic.lean | 21 ++++++++++++--------- 2 files changed, 13 insertions(+), 10 deletions(-) diff --git a/Mathlib/SetTheory/Cardinal/Aleph.lean b/Mathlib/SetTheory/Cardinal/Aleph.lean index ca79284dac29a..307901b8fdca5 100644 --- a/Mathlib/SetTheory/Cardinal/Aleph.lean +++ b/Mathlib/SetTheory/Cardinal/Aleph.lean @@ -275,7 +275,7 @@ theorem ord_preAleph (o : Ordinal) : (preAleph o).ord = preOmega o := by rw [← o.card_preOmega, (isInitial_preOmega o).ord_card] @[simp] -theorem type_cardinal : @type Cardinal (· < ·) _ = Ordinal.univ.{u, u + 1} := by +theorem type_cardinal : typeLT Cardinal = Ordinal.univ.{u, u + 1} := by rw [Ordinal.univ_id] exact Quotient.sound ⟨preAleph.symm.toRelIsoLT⟩ diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index edf1824a8eb39..1693d12052eee 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -128,6 +128,9 @@ namespace Ordinal def type (r : α → α → Prop) [wo : IsWellOrder α r] : Ordinal := ⟦⟨α, r, wo⟩⟧ +/-- `typeLT α` is an abbreviation for the order type of the `<` relation of `α`. -/ +scoped notation "typeLT " α:70 => @Ordinal.type α (· < ·) inferInstance + instance zero : Zero Ordinal := ⟨type <| @EmptyRelation PEmpty⟩ @@ -147,11 +150,11 @@ theorem type_def' (w : WellOrder) : ⟦w⟧ = type w.r := rfl theorem type_def (r) [wo : IsWellOrder α r] : (⟦⟨α, r, wo⟩⟧ : Ordinal) = type r := rfl @[simp] -theorem type_toType (o : Ordinal) : type (α := o.toType) (· < ·) = o := +theorem type_toType (o : Ordinal) : typeLT o.toType = o := o.out_eq @[deprecated type_toType (since := "2024-10-22")] -theorem type_lt (o : Ordinal) : type (α := o.toType) (· < ·) = o := +theorem type_lt (o : Ordinal) : typeLT o.toType = o := o.out_eq @[deprecated type_toType (since := "2024-08-26")] @@ -759,14 +762,14 @@ theorem lt_lift_iff {a : Ordinal.{u}} {b : Ordinal.{max u v}} : /-- `ω` is the first infinite ordinal, defined as the order type of `ℕ`. -/ def omega0 : Ordinal.{u} := - lift <| @type ℕ (· < ·) _ + lift (typeLT ℕ) @[inherit_doc] scoped notation "ω" => Ordinal.omega0 /-- Note that the presence of this lemma makes `simp [omega0]` form a loop. -/ @[simp] -theorem type_nat_lt : @type ℕ (· < ·) _ = ω := +theorem type_nat_lt : typeLT ℕ = ω := (lift_id _).symm @[simp] @@ -973,9 +976,9 @@ theorem le_enum_succ {o : Ordinal} (a : (succ o).toType) : of `Ordinal.{v}` (when `u < v`). It is an inaccessible cardinal. -/ @[pp_with_univ, nolint checkUnivs] def univ : Ordinal.{max (u + 1) v} := - lift.{v, u + 1} (@type Ordinal (· < ·) _) + lift.{v, u + 1} (typeLT Ordinal) -theorem univ_id : univ.{u, u + 1} = @type Ordinal (· < ·) _ := +theorem univ_id : univ.{u, u + 1} = typeLT Ordinal := lift_id _ @[simp] @@ -1039,12 +1042,12 @@ set_option linter.deprecated false in theorem lift.principalSeg_top : (lift.principalSeg.{u, v}).top = univ.{u, v} := rfl -theorem liftPrincipalSeg_top' : liftPrincipalSeg.{u, u + 1}.top = @type Ordinal (· < ·) _ := by +theorem liftPrincipalSeg_top' : liftPrincipalSeg.{u, u + 1}.top = typeLT Ordinal := by simp only [liftPrincipalSeg_top, univ_id] set_option linter.deprecated false in @[deprecated liftPrincipalSeg_top (since := "2024-09-21")] -theorem lift.principalSeg_top' : lift.principalSeg.{u, u + 1}.top = @type Ordinal (· < ·) _ := by +theorem lift.principalSeg_top' : lift.principalSeg.{u, u + 1}.top = typeLT Ordinal := by simp only [lift.principalSeg_top, univ_id] end Ordinal @@ -1426,7 +1429,7 @@ theorem card_eq_ofNat {o} {n : ℕ} [n.AtLeastTwo] : theorem type_fintype (r : α → α → Prop) [IsWellOrder α r] [Fintype α] : type r = Fintype.card α := by rw [← card_eq_nat, card_type, mk_fintype] -theorem type_fin (n : ℕ) : @type (Fin n) (· < ·) _ = n := by simp +theorem type_fin (n : ℕ) : typeLT (Fin n) = n := by simp end Ordinal From de7196a7d290d6269d4d5f1130a3d0a8c29414ce Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Mon, 2 Dec 2024 16:40:49 +0000 Subject: [PATCH 478/829] chore: turn `erw` into `rw` for free (#19689) The `erw` linter from #17638 caught these `erw`s that can become `rw` without breaking the proofs. --- Mathlib/AlgebraicGeometry/AffineSpace.lean | 2 +- Mathlib/Combinatorics/Enumerative/IncidenceAlgebra.lean | 4 ++-- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/AffineSpace.lean b/Mathlib/AlgebraicGeometry/AffineSpace.lean index 805e43c4928d2..7751d57d64eda 100644 --- a/Mathlib/AlgebraicGeometry/AffineSpace.lean +++ b/Mathlib/AlgebraicGeometry/AffineSpace.lean @@ -297,7 +297,7 @@ lemma map_comp {S S' S'' : Scheme} (f : S ⟶ S') (g : S' ⟶ S'') : · simp · simp only [TopologicalSpace.Opens.map_top, Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, Scheme.comp_app, CommRingCat.comp_apply] - erw [map_appTop_coord, map_appTop_coord, map_appTop_coord] + rw [map_appTop_coord, map_appTop_coord, map_appTop_coord] lemma map_Spec_map {R S : CommRingCat.{max u v}} (φ : R ⟶ S) : map n (Spec.map φ) = diff --git a/Mathlib/Combinatorics/Enumerative/IncidenceAlgebra.lean b/Mathlib/Combinatorics/Enumerative/IncidenceAlgebra.lean index 08a4f0b16172f..a049787977d46 100644 --- a/Mathlib/Combinatorics/Enumerative/IncidenceAlgebra.lean +++ b/Mathlib/Combinatorics/Enumerative/IncidenceAlgebra.lean @@ -554,8 +554,8 @@ lemma moebius_inversion_top (f g : α → 𝕜) (h : ∀ x, g x = ∑ y ∈ Ici rw [zeta_apply, if_pos (mem_Ici.mp ‹_›), one_mul] _ = ∑ y ∈ Ici x, ∑ z ∈ Ici y, mu 𝕜 x y * zeta 𝕜 y z * f z := by simp [mul_sum] _ = ∑ z ∈ Ici x, ∑ y ∈ Icc x z, mu 𝕜 x y * zeta 𝕜 y z * f z := by - erw [sum_sigma' (Ici x) fun y ↦ Ici y] - erw [sum_sigma' (Ici x) fun z ↦ Icc x z] + rw [sum_sigma' (Ici x) fun y ↦ Ici y] + rw [sum_sigma' (Ici x) fun z ↦ Icc x z] simp only [mul_boole, MulZeroClass.zero_mul, ite_mul, zeta_apply] apply sum_nbij' (fun ⟨a, b⟩ ↦ ⟨b, a⟩) (fun ⟨a, b⟩ ↦ ⟨b, a⟩) <;> aesop (add simp mul_assoc) (add unsafe le_trans) From c46519420dc1f5de05f5074f5df19126a32b9f8d Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 2 Dec 2024 17:27:25 +0000 Subject: [PATCH 479/829] feat: notation for vectors in Euclidean space (#17732) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This adds `!₂[1, 2, 3]` notation for `(WithLp.equiv 2 (∀ _ : Fin 3, _)).symm ![1, 2, 3]`, which is a term of type `EuclideanSpace 𝕜 (Fin 3)`; and more generally for other Lp norms, using the `subscript` parser. This reduces the temptation to write the type-incorrect `![1, 2, 3]` that has the wrong norm. The (parser for the) `subscript` parser relies on `initialize`rs running, and so a missing `Lean.enableInitializersExecution` has to be inserted in `checkYaml.lean`. It has already been inserted in the necessary places in upstream repos. [Notation Zulip poll](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Notation.20for.20vectors.20in.20euclidean.20space/near/476796192). Co-authored-by: Eric Wieser --- Mathlib/Analysis/InnerProductSpace/PiL2.lean | 39 +++++++++++++++++++- Mathlib/Analysis/Quaternion.lean | 2 +- MathlibTest/EuclideanSpace.lean | 33 +++++++++++++++++ scripts/check-yaml.lean | 1 + scripts/noshake.json | 1 + 5 files changed, 73 insertions(+), 3 deletions(-) create mode 100644 MathlibTest/EuclideanSpace.lean diff --git a/Mathlib/Analysis/InnerProductSpace/PiL2.lean b/Mathlib/Analysis/InnerProductSpace/PiL2.lean index 227dd212a2354..3469e47591c82 100644 --- a/Mathlib/Analysis/InnerProductSpace/PiL2.lean +++ b/Mathlib/Analysis/InnerProductSpace/PiL2.lean @@ -7,6 +7,7 @@ import Mathlib.Analysis.InnerProductSpace.Projection import Mathlib.Analysis.Normed.Lp.PiLp import Mathlib.LinearAlgebra.FiniteDimensional import Mathlib.LinearAlgebra.UnitaryGroup +import Mathlib.Util.Superscript /-! # `L²` inner product space structure on finite products of inner product spaces @@ -30,7 +31,8 @@ the last section, various properties of matrices are explored. - `EuclideanSpace 𝕜 n`: defined to be `PiLp 2 (n → 𝕜)` for any `Fintype n`, i.e., the space from functions to `n` to `𝕜` with the `L²` norm. We register several instances on it (notably - that it is a finite-dimensional inner product space). + that it is a finite-dimensional inner product space), and provide a `!ₚ[]` notation (for numeric + subscripts like `₂`) for the case when the indexing type is `Fin n`. - `OrthonormalBasis 𝕜 ι`: defined to be an isometry to Euclidean space from a given finite-dimensional inner product space, `E ≃ₗᵢ[𝕜] EuclideanSpace 𝕜 ι`. @@ -95,10 +97,43 @@ theorem PiLp.inner_apply {ι : Type*} [Fintype ι] {f : ι → Type*} [∀ i, No rfl /-- The standard real/complex Euclidean space, functions on a finite type. For an `n`-dimensional -space use `EuclideanSpace 𝕜 (Fin n)`. -/ +space use `EuclideanSpace 𝕜 (Fin n)`. + +For the case when `n = Fin _`, there is `!₂[x, y, ...]` notation for building elements of this type, +analogous to `![x, y, ...]` notation. -/ abbrev EuclideanSpace (𝕜 : Type*) (n : Type*) : Type _ := PiLp 2 fun _ : n => 𝕜 +section Notation +open Lean Meta Elab Term Macro TSyntax PrettyPrinter.Delaborator SubExpr + +/-- Notation for vectors in Lp space. `!₂[x, y, ...]` is a shorthand for +`(WithLp.equiv 2 _ _).symm ![x, y, ...]`, of type `EuclideanSpace _ (Fin _)`. + +This also works for other subscripts. -/ +syntax (name := PiLp.vecNotation) "!" noWs subscript(term) noWs "[" term,* "]" : term +macro_rules | `(!$p:subscript[$e:term,*]) => do + -- override the `Fin n.succ` to a literal + let n := e.getElems.size + `((WithLp.equiv $p <| ∀ _ : Fin $(quote n), _).symm ![$e,*]) + +set_option trace.debug true in +/-- Unexpander for the `!₂[x, y, ...]` notation. -/ +@[delab app.DFunLike.coe] +def EuclideanSpace.delabVecNotation : Delab := + whenNotPPOption getPPExplicit <| whenPPOption getPPNotation <| withOverApp 6 do + -- check that the `(WithLp.equiv _ _).symm` is present + let p : Term ← withAppFn <| withAppArg do + let_expr Equiv.symm _ _ e := ← getExpr | failure + let_expr WithLp.equiv _ _ := e | failure + withNaryArg 2 <| withNaryArg 0 <| delab + -- to be conservative, only allow subscripts which are numerals + guard <| p matches `($_:num) + let `(![$elems,*]) := ← withAppArg delab | failure + `(!$p[$elems,*]) + +end Notation + theorem EuclideanSpace.nnnorm_eq {𝕜 : Type*} [RCLike 𝕜] {n : Type*} [Fintype n] (x : EuclideanSpace 𝕜 n) : ‖x‖₊ = NNReal.sqrt (∑ i, ‖x i‖₊ ^ 2) := PiLp.nnnorm_eq_of_L2 x diff --git a/Mathlib/Analysis/Quaternion.lean b/Mathlib/Analysis/Quaternion.lean index ff600ebb0aa50..9f75be8137503 100644 --- a/Mathlib/Analysis/Quaternion.lean +++ b/Mathlib/Analysis/Quaternion.lean @@ -165,7 +165,7 @@ theorem norm_piLp_equiv_symm_equivTuple (x : ℍ) : noncomputable def linearIsometryEquivTuple : ℍ ≃ₗᵢ[ℝ] EuclideanSpace ℝ (Fin 4) := { (QuaternionAlgebra.linearEquivTuple (-1 : ℝ) (-1 : ℝ)).trans (WithLp.linearEquiv 2 ℝ (Fin 4 → ℝ)).symm with - toFun := fun a => (WithLp.equiv _ (Fin 4 → _)).symm ![a.1, a.2, a.3, a.4] + toFun := fun a => !₂[a.1, a.2, a.3, a.4] invFun := fun a => ⟨a 0, a 1, a 2, a 3⟩ norm_map' := norm_piLp_equiv_symm_equivTuple } diff --git a/MathlibTest/EuclideanSpace.lean b/MathlibTest/EuclideanSpace.lean new file mode 100644 index 0000000000000..662b91d2a0251 --- /dev/null +++ b/MathlibTest/EuclideanSpace.lean @@ -0,0 +1,33 @@ +import Mathlib.Analysis.InnerProductSpace.PiL2 + +#guard_expr !₂[] = (WithLp.equiv 2 (∀ _ : Fin 0, _)).symm ![] +#guard_expr !₂[1, 2, 3] = (WithLp.equiv 2 (∀ _ : Fin 3, ℕ)).symm ![1, 2, 3] +#guard_expr !₁[1, 2, (3 : ℝ)] = (WithLp.equiv 1 (∀ _ : Fin 3, ℝ)).symm ![1, 2, 3] + +section delaborator + +/-- info: !₂[1, 2, 3] : WithLp 2 (Fin 3 → ℕ) -/ +#guard_msgs in +#check !₂[1, 2, 3] + +set_option pp.mvars false in +/-- info: !₀[] : WithLp 0 (Fin 0 → ?_) -/ +#guard_msgs in +#check !₀[] + +section var +variable {p : ENNReal} +/-- info: (WithLp.equiv p (Fin 3 → ℕ)).symm ![1, 2, 3] : WithLp p (Fin 3 → ℕ) -/ +#guard_msgs in#check !ₚ[1, 2, 3] +end var + +section tombstoned_var +/- Here the `p` in the subscript is shadowed by a later p; so even if we do +make the delaborator less conservative, it should not fire here since `✝` cannot +be subscripted. -/ +variable {p : ENNReal} {x} (hx : x = !ₚ[1, 2, 3]) (p : True) +/-- info: hx : x = (WithLp.equiv p✝ (Fin 3 → ℕ)).symm ![1, 2, 3] -/ +#guard_msgs in #check hx +end tombstoned_var + +end delaborator diff --git a/scripts/check-yaml.lean b/scripts/check-yaml.lean index 731c46a1cabab..5d785ad1a5928 100644 --- a/scripts/check-yaml.lean +++ b/scripts/check-yaml.lean @@ -40,6 +40,7 @@ def processDb (decls : ConstMap) : String → IO Bool return false unsafe def main : IO Unit := do + Lean.enableInitializersExecution CoreM.withImportModules #[`Mathlib, `Archive] (searchPath := compile_time_search_path%) (trustLevel := 1024) do let decls := (← getEnv).constants diff --git a/scripts/noshake.json b/scripts/noshake.json index 3126619078308..acb992ed41e28 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -424,6 +424,7 @@ ["Mathlib.Tactic.CategoryTheory.Bicategory.Basic"], "Mathlib.Analysis.Normed.Operator.LinearIsometry": ["Mathlib.Algebra.Star.Basic"], + "Mathlib.Analysis.InnerProductSpace.PiL2": ["Mathlib.Util.Superscript"], "Mathlib.Analysis.InnerProductSpace.Basic": ["Mathlib.Algebra.Module.LinearMap.Basic"], "Mathlib.Analysis.Distribution.SchwartzSpace": ["Mathlib.Tactic.MoveAdd"], From 8b7b65fece8c0c914cbabf382c4e985cfacc7abf Mon Sep 17 00:00:00 2001 From: Seewoo Lee Date: Mon, 2 Dec 2024 17:40:40 +0000 Subject: [PATCH 480/829] feat: `isCoprime_mul_units_left`, `isCoprime_mul_units_right` (#19133) Co-authored-by: Seewoo Lee <49933279+seewoo5@users.noreply.github.com> --- Mathlib/RingTheory/Coprime/Basic.lean | 24 ++++++++++++++++++------ 1 file changed, 18 insertions(+), 6 deletions(-) diff --git a/Mathlib/RingTheory/Coprime/Basic.lean b/Mathlib/RingTheory/Coprime/Basic.lean index 851c87915e7a1..0e608f5a7df62 100644 --- a/Mathlib/RingTheory/Coprime/Basic.lean +++ b/Mathlib/RingTheory/Coprime/Basic.lean @@ -244,7 +244,7 @@ end ScalarTower section CommSemiringUnit -variable {R : Type*} [CommSemiring R] {x : R} +variable {R : Type*} [CommSemiring R] {x u v : R} theorem isCoprime_mul_unit_left_left (hu : IsUnit x) (y z : R) : IsCoprime (x * y) z ↔ IsCoprime y z := @@ -256,10 +256,6 @@ theorem isCoprime_mul_unit_left_right (hu : IsUnit x) (y z : R) : let ⟨u, hu⟩ := hu hu ▸ isCoprime_group_smul_right u y z -theorem isCoprime_mul_unit_left (hu : IsUnit x) (y z : R) : - IsCoprime (x * y) (x * z) ↔ IsCoprime y z := - (isCoprime_mul_unit_left_left hu y (x * z)).trans (isCoprime_mul_unit_left_right hu y z) - theorem isCoprime_mul_unit_right_left (hu : IsUnit x) (y z : R) : IsCoprime (y * x) z ↔ IsCoprime y z := mul_comm x y ▸ isCoprime_mul_unit_left_left hu y z @@ -268,9 +264,25 @@ theorem isCoprime_mul_unit_right_right (hu : IsUnit x) (y z : R) : IsCoprime y (z * x) ↔ IsCoprime y z := mul_comm x z ▸ isCoprime_mul_unit_left_right hu y z +theorem isCoprime_mul_units_left (hu : IsUnit u) (hv : IsUnit v) (y z : R) : + IsCoprime (u * y) (v * z) ↔ IsCoprime y z := + Iff.trans + (isCoprime_mul_unit_left_left hu _ _) + (isCoprime_mul_unit_left_right hv _ _) + +theorem isCoprime_mul_units_right (hu : IsUnit u) (hv : IsUnit v) (y z : R) : + IsCoprime (y * u) (z * v) ↔ IsCoprime y z := + Iff.trans + (isCoprime_mul_unit_right_left hu _ _) + (isCoprime_mul_unit_right_right hv _ _) + +theorem isCoprime_mul_unit_left (hu : IsUnit x) (y z : R) : + IsCoprime (x * y) (x * z) ↔ IsCoprime y z := + isCoprime_mul_units_left hu hu _ _ + theorem isCoprime_mul_unit_right (hu : IsUnit x) (y z : R) : IsCoprime (y * x) (z * x) ↔ IsCoprime y z := - (isCoprime_mul_unit_right_left hu y (z * x)).trans (isCoprime_mul_unit_right_right hu y z) + isCoprime_mul_units_right hu hu _ _ end CommSemiringUnit From 561e050467d2caa2e4ea5d81ca62600bf268b747 Mon Sep 17 00:00:00 2001 From: FMLJohn <417963616@qq.com> Date: Mon, 2 Dec 2024 18:04:18 +0000 Subject: [PATCH 481/829] =?UTF-8?q?feat(RingTheory/Polynomial/HilbertPoly)?= =?UTF-8?q?:=20the=20definition=20and=20key=20property=20of=20`Polynomial.?= =?UTF-8?q?hilbertPoly=20p=20d`=20for=20`p=20:=20F[X]`=20and=20`d=20:=20?= =?UTF-8?q?=E2=84=95`,=20where=20`F`=20is=20a=20field.=20(#19303)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Given any field `F`, polynomial `p : F[X]` and natural number `d`, we have defined `Polynomial.hilbertPoly p d : F[X]`. If `F` is of characteristic zero, then for any large enough `n : ℕ`, `PowerSeries.coeff F n (p * (invOneSubPow F d))` equals `(hilbertPoly p d).eval (n : F)` (see `Polynomial.coeff_mul_invOneSubPow_eq_hilbertPoly_eval`). --- Mathlib.lean | 1 + Mathlib/Data/Nat/Factorial/BigOperators.lean | 4 + .../RingTheory/Polynomial/HilbertPoly.lean | 130 ++++++++++++++++++ Mathlib/RingTheory/Polynomial/Pochhammer.lean | 10 ++ Mathlib/RingTheory/PowerSeries/WellKnown.lean | 5 +- 5 files changed, 147 insertions(+), 3 deletions(-) create mode 100644 Mathlib/RingTheory/Polynomial/HilbertPoly.lean diff --git a/Mathlib.lean b/Mathlib.lean index 380263015c62c..ac692cd27c721 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4481,6 +4481,7 @@ import Mathlib.RingTheory.Polynomial.Eisenstein.IsIntegral import Mathlib.RingTheory.Polynomial.GaussLemma import Mathlib.RingTheory.Polynomial.Hermite.Basic import Mathlib.RingTheory.Polynomial.Hermite.Gaussian +import Mathlib.RingTheory.Polynomial.HilbertPoly import Mathlib.RingTheory.Polynomial.Ideal import Mathlib.RingTheory.Polynomial.IntegralNormalization import Mathlib.RingTheory.Polynomial.IrreducibleRing diff --git a/Mathlib/Data/Nat/Factorial/BigOperators.lean b/Mathlib/Data/Nat/Factorial/BigOperators.lean index 56a55af7e4af1..606f5b7c49b2d 100644 --- a/Mathlib/Data/Nat/Factorial/BigOperators.lean +++ b/Mathlib/Data/Nat/Factorial/BigOperators.lean @@ -33,6 +33,10 @@ theorem prod_factorial_dvd_factorial_sum : (∏ i ∈ s, (f i)!) ∣ (∑ i ∈ · rw [prod_cons, Finset.sum_cons] exact (mul_dvd_mul_left _ ih).trans (Nat.factorial_mul_factorial_dvd_factorial_add _ _) +theorem ascFactorial_eq_prod_range (n : ℕ) : ∀ k, n.ascFactorial k = ∏ i ∈ range k, (n + i) + | 0 => rfl + | k + 1 => by rw [ascFactorial, prod_range_succ, mul_comm, ascFactorial_eq_prod_range n k] + theorem descFactorial_eq_prod_range (n : ℕ) : ∀ k, n.descFactorial k = ∏ i ∈ range k, (n - i) | 0 => rfl | k + 1 => by rw [descFactorial, prod_range_succ, mul_comm, descFactorial_eq_prod_range n k] diff --git a/Mathlib/RingTheory/Polynomial/HilbertPoly.lean b/Mathlib/RingTheory/Polynomial/HilbertPoly.lean new file mode 100644 index 0000000000000..302a9660f7465 --- /dev/null +++ b/Mathlib/RingTheory/Polynomial/HilbertPoly.lean @@ -0,0 +1,130 @@ +/- +Copyright (c) 2024 Fangming Li. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Fangming Li, Jujian Zhang +-/ +import Mathlib.Algebra.Polynomial.AlgebraMap +import Mathlib.Algebra.Polynomial.Eval.SMul +import Mathlib.RingTheory.Polynomial.Pochhammer +import Mathlib.RingTheory.PowerSeries.WellKnown +import Mathlib.Tactic.FieldSimp + +/-! +# Hilbert polynomials + +In this file, we formalise the following statement: if `F` is a field with characteristic `0`, then +given any `p : F[X]` and `d : ℕ`, there exists some `h : F[X]` such that for any large enough +`n : ℕ`, `h(n)` is equal to the coefficient of `Xⁿ` in the power series expansion of `p/(1 - X)ᵈ`. +This `h` is unique and is denoted as `Polynomial.hilbertPoly p d`. + +For example, given `d : ℕ`, the power series expansion of `1/(1-X)ᵈ⁺¹` in `F[X]` is +`Σₙ ((d + n).choose d)Xⁿ`, which equals `Σₙ ((n + 1)···(n + d)/d!)Xⁿ` and hence +`Polynomial.hilbertPoly (1 : F[X]) d` is the polynomial `(n + 1)···(n + d)/d!`. Note that +if `d! = 0` in `F`, then the polynomial `(n + 1)···(n + d)/d!` no longer works, so we do not +want the characteristic of `F` to be divisible by `d!`. As `Polynomial.hilbertPoly` may take +any `p : F[X]` and `d : ℕ` as its inputs, it is necessary for us to assume that `CharZero F`. + +## Main definitions + +* `Polynomial.hilbertPoly p d`. Given a field `F`, a polynomial `p : F[X]` and a natural number `d`, + if `F` is of characteristic `0`, then `Polynomial.hilbertPoly p d : F[X]` is the polynomial whose + value at `n` equals the coefficient of `Xⁿ` in the power series expansion of `p/(1 - X)ᵈ`. + +## TODO + +* Hilbert polynomials of finitely generated graded modules over Noetherian rings. +-/ + +open Nat PowerSeries + +variable (F : Type*) [Field F] + +namespace Polynomial + +/-- +For any field `F` and natural numbers `d` and `k`, `Polynomial.preHilbertPoly F d k` +is defined as `(d.factorial : F)⁻¹ • ((ascPochhammer F d).comp (X - (C (k : F)) + 1))`. +This is the most basic form of Hilbert polynomials. `Polynomial.preHilbertPoly ℚ d 0` +is exactly the Hilbert polynomial of the polynomial ring `ℚ[X_0,...,X_d]` viewed as +a graded module over itself. In fact, `Polynomial.preHilbertPoly F d k` is the +same as `Polynomial.hilbertPoly ((X : F[X]) ^ k) (d + 1)` for any field `F` and +`d k : ℕ` (see the lemma `Polynomial.hilbertPoly_X_pow_succ`). See also the lemma +`Polynomial.preHilbertPoly_eq_choose_sub_add`, which states that if `CharZero F`, +then for any `d k n : ℕ` with `k ≤ n`, `(Polynomial.preHilbertPoly F d k).eval (n : F)` +equals `(n - k + d).choose d`. +-/ +noncomputable def preHilbertPoly (d k : ℕ) : F[X] := + (d.factorial : F)⁻¹ • ((ascPochhammer F d).comp (Polynomial.X - (C (k : F)) + 1)) + +lemma preHilbertPoly_eq_choose_sub_add [CharZero F] (d : ℕ) {k n : ℕ} (hkn : k ≤ n): + (preHilbertPoly F d k).eval (n : F) = (n - k + d).choose d := by + have : ((d ! : ℕ) : F) ≠ 0 := by norm_cast; positivity + calc + _ = (↑d !)⁻¹ * eval (↑(n - k + 1)) (ascPochhammer F d) := by simp [cast_sub hkn, preHilbertPoly] + _ = (n - k + d).choose d := by + rw [ascPochhammer_nat_eq_natCast_ascFactorial]; + field_simp [ascFactorial_eq_factorial_mul_choose] + +variable {F} + +/-- +`Polynomial.hilbertPoly p 0 = 0`; for any `d : ℕ`, `Polynomial.hilbertPoly p (d + 1)` +is defined as `∑ i in p.support, (p.coeff i) • Polynomial.preHilbertPoly F d i`. If +`M` is a graded module whose Poincaré series can be written as `p(X)/(1 - X)ᵈ` for some +`p : ℚ[X]` with integer coefficients, then `Polynomial.hilbertPoly p d` is the Hilbert +polynomial of `M`. See also `Polynomial.coeff_mul_invOneSubPow_eq_hilbertPoly_eval`, +which says that `PowerSeries.coeff F n (p * (PowerSeries.invOneSubPow F d))` equals +`(Polynomial.hilbertPoly p d).eval (n : F)` for any large enough `n : ℕ`. +-/ +noncomputable def hilbertPoly (p : F[X]) : (d : ℕ) → F[X] + | 0 => 0 + | d + 1 => ∑ i in p.support, (p.coeff i) • preHilbertPoly F d i + +variable (F) in +lemma hilbertPoly_zero_nat (d : ℕ) : hilbertPoly (0 : F[X]) d = 0 := by + delta hilbertPoly; induction d with + | zero => simp only + | succ d _ => simp only [coeff_zero, zero_smul, Finset.sum_const_zero] + +lemma hilbertPoly_poly_zero (p : F[X]) : hilbertPoly p 0 = 0 := rfl + +lemma hilbertPoly_poly_succ (p : F[X]) (d : ℕ) : + hilbertPoly p (d + 1) = ∑ i in p.support, (p.coeff i) • preHilbertPoly F d i := rfl + +variable (F) in +lemma hilbertPoly_X_pow_succ (d k : ℕ) : + hilbertPoly ((X : F[X]) ^ k) (d + 1) = preHilbertPoly F d k := by + delta hilbertPoly; simp + +/-- +The key property of Hilbert polynomials. If `F` is a field with characteristic `0`, `p : F[X]` and +`d : ℕ`, then for any large enough `n : ℕ`, `(Polynomial.hilbertPoly p d).eval (n : F)` equals the +coefficient of `Xⁿ` in the power series expansion of `p/(1 - X)ᵈ`. +-/ +theorem coeff_mul_invOneSubPow_eq_hilbertPoly_eval + [CharZero F] {p : F[X]} (d : ℕ) {n : ℕ} (hn : p.natDegree < n) : + PowerSeries.coeff F n (p * (invOneSubPow F d)) = (hilbertPoly p d).eval (n : F) := by + delta hilbertPoly; induction d with + | zero => simp only [invOneSubPow_zero, Units.val_one, mul_one, coeff_coe, eval_zero] + exact coeff_eq_zero_of_natDegree_lt hn + | succ d hd => + simp only [eval_finset_sum, eval_smul, smul_eq_mul] + rw [← Finset.sum_coe_sort] + have h_le (i : p.support) : (i : ℕ) ≤ n := + le_trans (le_natDegree_of_ne_zero <| mem_support_iff.1 i.2) hn.le + have h (i : p.support) : eval ↑n (preHilbertPoly F d ↑i) = (n + d - ↑i).choose d := by + rw [preHilbertPoly_eq_choose_sub_add _ _ (h_le i), Nat.sub_add_comm (h_le i)] + simp_rw [h] + rw [Finset.sum_coe_sort _ (fun x => (p.coeff ↑x) * (_ + d - ↑x).choose _), + PowerSeries.coeff_mul, Finset.Nat.sum_antidiagonal_eq_sum_range_succ_mk, + invOneSubPow_val_eq_mk_sub_one_add_choose_of_pos _ _ (zero_lt_succ d)] + simp only [coeff_coe, coeff_mk] + symm + refine Finset.sum_subset_zero_on_sdiff (fun s hs ↦ ?_) (fun x hx ↦ ?_) (fun x hx ↦ ?_) + · rw [Finset.mem_range_succ_iff] + exact h_le ⟨s, hs⟩ + · simp only [Finset.mem_sdiff, mem_support_iff, not_not] at hx + rw [hx.2, zero_mul] + · rw [add_comm, Nat.add_sub_assoc (h_le ⟨x, hx⟩), succ_eq_add_one, add_tsub_cancel_right] + +end Polynomial diff --git a/Mathlib/RingTheory/Polynomial/Pochhammer.lean b/Mathlib/RingTheory/Polynomial/Pochhammer.lean index d6872afb27cbc..29176297e9018 100644 --- a/Mathlib/RingTheory/Polynomial/Pochhammer.lean +++ b/Mathlib/RingTheory/Polynomial/Pochhammer.lean @@ -153,10 +153,20 @@ theorem ascPochhammer_nat_eq_ascFactorial (n : ℕ) : rw [ascPochhammer_succ_right, eval_mul, ascPochhammer_nat_eq_ascFactorial n t, eval_add, eval_X, eval_natCast, Nat.cast_id, Nat.ascFactorial_succ, mul_comm] +theorem ascPochhammer_nat_eq_natCast_ascFactorial (S : Type*) [Semiring S] (n k : ℕ) : + (ascPochhammer S k).eval (n : S) = n.ascFactorial k := by + norm_cast + rw [ascPochhammer_nat_eq_ascFactorial] + theorem ascPochhammer_nat_eq_descFactorial (a b : ℕ) : (ascPochhammer ℕ b).eval a = (a + b - 1).descFactorial b := by rw [ascPochhammer_nat_eq_ascFactorial, Nat.add_descFactorial_eq_ascFactorial'] +theorem ascPochhammer_nat_eq_natCast_descFactorial (S : Type*) [Semiring S] (a b : ℕ) : + (ascPochhammer S b).eval (a : S) = (a + b - 1).descFactorial b := by + norm_cast + rw [ascPochhammer_nat_eq_descFactorial] + @[simp] theorem ascPochhammer_natDegree (n : ℕ) [NoZeroDivisors S] [Nontrivial S] : (ascPochhammer S n).natDegree = n := by diff --git a/Mathlib/RingTheory/PowerSeries/WellKnown.lean b/Mathlib/RingTheory/PowerSeries/WellKnown.lean index ef67977a0a5d2..ca5bd0b442330 100644 --- a/Mathlib/RingTheory/PowerSeries/WellKnown.lean +++ b/Mathlib/RingTheory/PowerSeries/WellKnown.lean @@ -151,10 +151,9 @@ theorem invOneSubPow_inv_eq_one_sub_pow : | zero => exact Eq.symm <| pow_zero _ | succ d => rfl -theorem invOneSubPow_inv_eq_one_of_eq_zero (h : d = 0) : - (invOneSubPow S d).inv = 1 := by +theorem invOneSubPow_inv_zero_eq_one : (invOneSubPow S 0).inv = 1 := by delta invOneSubPow - simp only [h, Units.inv_eq_val_inv, inv_one, Units.val_one] + simp only [Units.inv_eq_val_inv, inv_one, Units.val_one] theorem mk_add_choose_mul_one_sub_pow_eq_one : (mk fun n ↦ Nat.choose (d + n) d : S⟦X⟧) * ((1 - X) ^ (d + 1)) = 1 := From 8a8d499eac0eb348953a126b83f44a0d90e14d4f Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 3 Dec 2024 03:19:18 +0000 Subject: [PATCH 482/829] chore: update username in maintainer list (#19702) --- README.md | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/README.md b/README.md index 6e5a49dbeafd5..5dd59489b323e 100644 --- a/README.md +++ b/README.md @@ -150,7 +150,7 @@ For a list containing more detailed information, see https://leanprover-communit * Patrick Massot (@patrickmassot): documentation, topology, geometry * Bhavik Mehta (@b-mehta): category theory, combinatorics * Kyle Miller (@kmill): combinatorics, tactics, metaprogramming -* Kim Morrison (@semorrison): category theory, tactics +* Kim Morrison (@kim-em): category theory, tactics * Oliver Nash (@ocfnash): algebra, geometry, topology * Joël Riou (@joelriou): category theory, homology, algebraic geometry * Damiano Testa (@adomani): algebra, algebraic geometry, number theory, tactics From 72b0941b7c0d3b285a99b48eab24cc2eeb673879 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 3 Dec 2024 05:50:37 +0000 Subject: [PATCH 483/829] chore: complete deprecation of `PartENat.card` (#19648) - [x] depends on: #19622 (perhaps unnecessarily) --- Mathlib.lean | 4 +- Mathlib/Algebra/Group/Subgroup/Finite.lean | 1 + .../Connectivity/WalkCounting.lean | 1 + Mathlib/Data/Finite/Card.lean | 15 --- Mathlib/Data/Matroid/Basic.lean | 1 + Mathlib/Data/Matroid/Closure.lean | 1 + Mathlib/Data/Matroid/IndepAxioms.lean | 2 + Mathlib/Deprecated/Cardinal/Continuum.lean | 25 ++++ Mathlib/Deprecated/Cardinal/Finite.lean | 121 ++++++++++++++++++ .../Cardinal/PartENat.lean | 0 Mathlib/GroupTheory/Order/Min.lean | 1 + Mathlib/GroupTheory/OrderOfElement.lean | 1 + Mathlib/NumberTheory/FLT/Three.lean | 1 - Mathlib/SetTheory/Cardinal/Aleph.lean | 7 +- Mathlib/SetTheory/Cardinal/Cofinality.lean | 1 + Mathlib/SetTheory/Cardinal/Continuum.lean | 4 +- Mathlib/SetTheory/Cardinal/Finite.lean | 103 +-------------- 17 files changed, 166 insertions(+), 123 deletions(-) create mode 100644 Mathlib/Deprecated/Cardinal/Continuum.lean create mode 100644 Mathlib/Deprecated/Cardinal/Finite.lean rename Mathlib/{SetTheory => Deprecated}/Cardinal/PartENat.lean (100%) diff --git a/Mathlib.lean b/Mathlib.lean index ac692cd27c721..09bc8fed5e198 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2925,6 +2925,9 @@ import Mathlib.Data.ZMod.Units import Mathlib.Deprecated.AlgebraClasses import Mathlib.Deprecated.Aliases import Mathlib.Deprecated.ByteArray +import Mathlib.Deprecated.Cardinal.Continuum +import Mathlib.Deprecated.Cardinal.Finite +import Mathlib.Deprecated.Cardinal.PartENat import Mathlib.Deprecated.Combinator import Mathlib.Deprecated.Equiv import Mathlib.Deprecated.Group @@ -4618,7 +4621,6 @@ import Mathlib.SetTheory.Cardinal.ENat import Mathlib.SetTheory.Cardinal.Finite import Mathlib.SetTheory.Cardinal.Finsupp import Mathlib.SetTheory.Cardinal.Free -import Mathlib.SetTheory.Cardinal.PartENat import Mathlib.SetTheory.Cardinal.SchroederBernstein import Mathlib.SetTheory.Cardinal.Subfield import Mathlib.SetTheory.Cardinal.ToNat diff --git a/Mathlib/Algebra/Group/Subgroup/Finite.lean b/Mathlib/Algebra/Group/Subgroup/Finite.lean index db72bb5f2f470..05d21ec805804 100644 --- a/Mathlib/Algebra/Group/Subgroup/Finite.lean +++ b/Mathlib/Algebra/Group/Subgroup/Finite.lean @@ -6,6 +6,7 @@ Authors: Kexing Ying import Mathlib.Algebra.Group.Subgroup.Basic import Mathlib.Algebra.Group.Submonoid.BigOperators import Mathlib.Data.Finite.Card +import Mathlib.Data.Set.Finite.Range /-! # Subgroups diff --git a/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean b/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean index 171bc81387364..559e58c4930b6 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.BigOperators.Ring.Nat import Mathlib.Combinatorics.SimpleGraph.Path import Mathlib.Combinatorics.SimpleGraph.Subgraph import Mathlib.SetTheory.Cardinal.Finite +import Mathlib.Data.Set.Finite.Lattice /-! # Counting walks of a given length diff --git a/Mathlib/Data/Finite/Card.lean b/Mathlib/Data/Finite/Card.lean index 4f038d6ecdef9..00700733c3af1 100644 --- a/Mathlib/Data/Finite/Card.lean +++ b/Mathlib/Data/Finite/Card.lean @@ -172,21 +172,6 @@ theorem card_eq_coe_natCard (α : Type*) [Finite α] : card α = Nat.card α := end ENat -namespace PartENat - -set_option linter.deprecated false in -@[deprecated ENat.card_eq_coe_natCard (since := "2024-11-30")] -theorem card_eq_coe_natCard (α : Type*) [Finite α] : card α = Nat.card α := by - unfold PartENat.card - apply symm - rw [Cardinal.natCast_eq_toPartENat_iff] - exact Finite.cast_card_eq_mk - - -@[deprecated (since := "2024-05-25")] alias card_eq_coe_nat_card := card_eq_coe_natCard - -end PartENat - namespace Set theorem card_union_le (s t : Set α) : Nat.card (↥(s ∪ t)) ≤ Nat.card s + Nat.card t := by diff --git a/Mathlib/Data/Matroid/Basic.lean b/Mathlib/Data/Matroid/Basic.lean index d3e21b40330bb..a3d760e295283 100644 --- a/Mathlib/Data/Matroid/Basic.lean +++ b/Mathlib/Data/Matroid/Basic.lean @@ -6,6 +6,7 @@ Authors: Peter Nelson import Mathlib.Data.Finite.Prod import Mathlib.Data.Matroid.Init import Mathlib.Data.Set.Card +import Mathlib.Data.Set.Finite.Powerset import Mathlib.Order.Minimal /-! diff --git a/Mathlib/Data/Matroid/Closure.lean b/Mathlib/Data/Matroid/Closure.lean index 2dbba090c7c1d..15b24edc449b6 100644 --- a/Mathlib/Data/Matroid/Closure.lean +++ b/Mathlib/Data/Matroid/Closure.lean @@ -5,6 +5,7 @@ Authors: Peter Nelson -/ import Mathlib.Data.Matroid.Restrict import Mathlib.Order.Closure +import Mathlib.Order.CompleteLatticeIntervals /-! # Matroid Closure diff --git a/Mathlib/Data/Matroid/IndepAxioms.lean b/Mathlib/Data/Matroid/IndepAxioms.lean index a1769ebc4fc2f..a18f6951b4618 100644 --- a/Mathlib/Data/Matroid/IndepAxioms.lean +++ b/Mathlib/Data/Matroid/IndepAxioms.lean @@ -3,6 +3,8 @@ Copyright (c) 2023 Peter Nelson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Peter Nelson -/ +import Mathlib.Order.Interval.Finset.Nat +import Mathlib.Data.Set.Finite.Lattice import Mathlib.Data.Matroid.Basic /-! diff --git a/Mathlib/Deprecated/Cardinal/Continuum.lean b/Mathlib/Deprecated/Cardinal/Continuum.lean new file mode 100644 index 0000000000000..42c2d813ef442 --- /dev/null +++ b/Mathlib/Deprecated/Cardinal/Continuum.lean @@ -0,0 +1,25 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Mario Carneiro, Floris van Doorn, Violeta Hernández Palacios +-/ +import Mathlib.SetTheory.Cardinal.Continuum +import Mathlib.Deprecated.Cardinal.PartENat + +/-! +# Deprecated material from Mathlib.SetTheory.Cardinal.Aleph and Mathlib.SetTheory.Cardinal.Continuum + +Moved here so we can reduce imports sooner. +-/ + +namespace Cardinal + +@[simp, deprecated aleph_toENat (since := "2024-12-01")] +theorem aleph_toPartENat (o : Ordinal) : toPartENat (ℵ_ o) = ⊤ := + toPartENat_apply_of_aleph0_le <| aleph0_le_aleph o + +@[simp, deprecated aleph_toENat (since := "2024-12-01")] +theorem continuum_toPartENat : toPartENat continuum = ⊤ := + toPartENat_apply_of_aleph0_le aleph0_le_continuum + +end Cardinal diff --git a/Mathlib/Deprecated/Cardinal/Finite.lean b/Mathlib/Deprecated/Cardinal/Finite.lean new file mode 100644 index 0000000000000..f28497050cfda --- /dev/null +++ b/Mathlib/Deprecated/Cardinal/Finite.lean @@ -0,0 +1,121 @@ +/- +Copyright (c) 2021 Aaron Anderson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Aaron Anderson +-/ +import Mathlib.Deprecated.Cardinal.PartENat +import Mathlib.SetTheory.Cardinal.Finite +import Mathlib.Data.Finite.Card + +/-! +# Deprecated material on `PartENat.card`. +-/ + +noncomputable section + +universe u v +variable {α : Type u} + +open Cardinal + +namespace PartENat + +/-- `PartENat.card α` is the cardinality of `α` as an extended natural number. + If `α` is infinite, `PartENat.card α = ⊤`. -/ +@[deprecated ENat.card (since := "2024-12-01")] +def card (α : Type*) : PartENat := + toPartENat (mk α) + +-- This rest of this file is about the deprecated `PartENat.card`. +set_option linter.deprecated false + +@[simp] +theorem card_eq_coe_fintype_card [Fintype α] : card α = Fintype.card α := + mk_toPartENat_eq_coe_card + +@[simp] +theorem card_eq_top_of_infinite [Infinite α] : card α = ⊤ := + mk_toPartENat_of_infinite + +@[simp] +theorem card_sum (α β : Type*) : + PartENat.card (α ⊕ β) = PartENat.card α + PartENat.card β := by + simp only [PartENat.card, Cardinal.mk_sum, map_add, Cardinal.toPartENat_lift] + +theorem card_congr {α : Type*} {β : Type*} (f : α ≃ β) : PartENat.card α = PartENat.card β := + Cardinal.toPartENat_congr f + +@[simp] lemma card_ulift (α : Type*) : card (ULift α) = card α := card_congr Equiv.ulift + +@[simp] lemma card_plift (α : Type*) : card (PLift α) = card α := card_congr Equiv.plift + +theorem card_image_of_injOn {α : Type u} {β : Type v} {f : α → β} {s : Set α} (h : Set.InjOn f s) : + card (f '' s) = card s := + card_congr (Equiv.Set.imageOfInjOn f s h).symm + +theorem card_image_of_injective {α : Type u} {β : Type v} (f : α → β) (s : Set α) + (h : Function.Injective f) : card (f '' s) = card s := card_image_of_injOn h.injOn + +-- Should I keep the 6 following lemmas ? +-- TODO: Add ofNat, zero, and one versions for simp confluence +@[simp] +theorem _root_.Cardinal.natCast_le_toPartENat_iff {n : ℕ} {c : Cardinal} : + ↑n ≤ toPartENat c ↔ ↑n ≤ c := by + rw [← toPartENat_natCast n, toPartENat_le_iff_of_le_aleph0 (le_of_lt (nat_lt_aleph0 n))] + +@[simp] +theorem _root_.Cardinal.toPartENat_le_natCast_iff {c : Cardinal} {n : ℕ} : + toPartENat c ≤ n ↔ c ≤ n := by + rw [← toPartENat_natCast n, toPartENat_le_iff_of_lt_aleph0 (nat_lt_aleph0 n)] + +@[simp] +theorem _root_.Cardinal.natCast_eq_toPartENat_iff {n : ℕ} {c : Cardinal} : + ↑n = toPartENat c ↔ ↑n = c := by + rw [le_antisymm_iff, le_antisymm_iff, Cardinal.toPartENat_le_natCast_iff, + Cardinal.natCast_le_toPartENat_iff] + +@[simp] +theorem _root_.Cardinal.toPartENat_eq_natCast_iff {c : Cardinal} {n : ℕ} : + Cardinal.toPartENat c = n ↔ c = n := by +rw [eq_comm, Cardinal.natCast_eq_toPartENat_iff, eq_comm] + +@[simp] +theorem _root_.Cardinal.natCast_lt_toPartENat_iff {n : ℕ} {c : Cardinal} : + ↑n < toPartENat c ↔ ↑n < c := by + simp only [← not_le, Cardinal.toPartENat_le_natCast_iff] + +@[simp] +theorem _root_.Cardinal.toPartENat_lt_natCast_iff {n : ℕ} {c : Cardinal} : + toPartENat c < ↑n ↔ c < ↑n := by + simp only [← not_le, Cardinal.natCast_le_toPartENat_iff] + +theorem card_eq_zero_iff_empty (α : Type*) : card α = 0 ↔ IsEmpty α := by + rw [← Cardinal.mk_eq_zero_iff] + conv_rhs => rw [← Nat.cast_zero] + simp only [← Cardinal.toPartENat_eq_natCast_iff] + simp only [PartENat.card, Nat.cast_zero] + +theorem card_le_one_iff_subsingleton (α : Type*) : card α ≤ 1 ↔ Subsingleton α := by + rw [← le_one_iff_subsingleton] + conv_rhs => rw [← Nat.cast_one] + rw [← Cardinal.toPartENat_le_natCast_iff] + simp only [PartENat.card, Nat.cast_one] + +theorem one_lt_card_iff_nontrivial (α : Type*) : 1 < card α ↔ Nontrivial α := by + rw [← Cardinal.one_lt_iff_nontrivial] + conv_rhs => rw [← Nat.cast_one] + rw [← natCast_lt_toPartENat_iff] + simp only [PartENat.card, Nat.cast_one] + +set_option linter.deprecated false in +@[deprecated ENat.card_eq_coe_natCard (since := "2024-11-30")] +theorem card_eq_coe_natCard (α : Type*) [Finite α] : card α = Nat.card α := by + unfold PartENat.card + apply symm + rw [Cardinal.natCast_eq_toPartENat_iff] + exact Finite.cast_card_eq_mk + + +@[deprecated (since := "2024-05-25")] alias card_eq_coe_nat_card := card_eq_coe_natCard + +end PartENat diff --git a/Mathlib/SetTheory/Cardinal/PartENat.lean b/Mathlib/Deprecated/Cardinal/PartENat.lean similarity index 100% rename from Mathlib/SetTheory/Cardinal/PartENat.lean rename to Mathlib/Deprecated/Cardinal/PartENat.lean diff --git a/Mathlib/GroupTheory/Order/Min.lean b/Mathlib/GroupTheory/Order/Min.lean index 3d1047b06f179..3b1ced396e120 100644 --- a/Mathlib/GroupTheory/Order/Min.lean +++ b/Mathlib/GroupTheory/Order/Min.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.GroupTheory.Torsion +import Mathlib.Data.ENat.Lattice /-! # Minimum order of an element diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index 564cc2da28dcb..492cc93e69891 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -10,6 +10,7 @@ import Mathlib.Algebra.Order.Group.Action import Mathlib.Algebra.Order.Ring.Abs import Mathlib.GroupTheory.Index import Mathlib.Order.Interval.Set.Infinite +import Mathlib.Tactic.Positivity /-! # Order of an element diff --git a/Mathlib/NumberTheory/FLT/Three.lean b/Mathlib/NumberTheory/FLT/Three.lean index 05013ac559c40..7f70138544592 100644 --- a/Mathlib/NumberTheory/FLT/Three.lean +++ b/Mathlib/NumberTheory/FLT/Three.lean @@ -321,7 +321,6 @@ section DecidableRel variable [NumberField K] [IsCyclotomicExtension {3} ℚ K] [DecidableRel fun (a b : 𝓞 K) ↦ a ∣ b] -open PartENat in /-- Given `S' : Solution'`, we have that `λ ^ 2` divides one amongst `S'.a + S'.b`, `S'.a + η * S'.b` and `S'.a + η ^ 2 * S'.b`. -/ lemma lambda_sq_dvd_or_dvd_or_dvd : diff --git a/Mathlib/SetTheory/Cardinal/Aleph.lean b/Mathlib/SetTheory/Cardinal/Aleph.lean index 307901b8fdca5..ae9eeee389945 100644 --- a/Mathlib/SetTheory/Cardinal/Aleph.lean +++ b/Mathlib/SetTheory/Cardinal/Aleph.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Floris van Doorn, Violeta Hernández Palacios -/ import Mathlib.Order.Bounded -import Mathlib.SetTheory.Cardinal.PartENat +import Mathlib.SetTheory.Cardinal.ToNat +import Mathlib.SetTheory.Cardinal.ENat import Mathlib.SetTheory.Ordinal.Enum /-! @@ -423,8 +424,8 @@ theorem aleph_toNat (o : Ordinal) : toNat (ℵ_ o) = 0 := toNat_apply_of_aleph0_le <| aleph0_le_aleph o @[simp] -theorem aleph_toPartENat (o : Ordinal) : toPartENat (ℵ_ o) = ⊤ := - toPartENat_apply_of_aleph0_le <| aleph0_le_aleph o +theorem aleph_toENat (o : Ordinal) : toENat (ℵ_ o) = ⊤ := + (toENat_eq_top.2 (aleph0_le_aleph o)) instance nonempty_toType_aleph (o : Ordinal) : Nonempty (ℵ_ o).ord.toType := by rw [toType_nonempty_iff_ne_zero, ← ord_zero] diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index 9a8363e186c9a..df29d51ac8b46 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -3,6 +3,7 @@ Copyright (c) 2017 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Floris van Doorn, Violeta Hernández Palacios -/ +import Mathlib.Data.Set.Finite.Lattice import Mathlib.SetTheory.Cardinal.Arithmetic import Mathlib.SetTheory.Ordinal.FixedPoint diff --git a/Mathlib/SetTheory/Cardinal/Continuum.lean b/Mathlib/SetTheory/Cardinal/Continuum.lean index b05977809225d..fd2ba0a5e9087 100644 --- a/Mathlib/SetTheory/Cardinal/Continuum.lean +++ b/Mathlib/SetTheory/Cardinal/Continuum.lean @@ -87,8 +87,8 @@ theorem continuum_toNat : toNat continuum = 0 := toNat_apply_of_aleph0_le aleph0_le_continuum @[simp] -theorem continuum_toPartENat : toPartENat continuum = ⊤ := - toPartENat_apply_of_aleph0_le aleph0_le_continuum +theorem continuum_toENat : toENat continuum = ⊤ := + (toENat_eq_top.2 aleph0_le_continuum) /-! ### Addition diff --git a/Mathlib/SetTheory/Cardinal/Finite.lean b/Mathlib/SetTheory/Cardinal/Finite.lean index 6498c5e18a81a..ed8bad925692c 100644 --- a/Mathlib/SetTheory/Cardinal/Finite.lean +++ b/Mathlib/SetTheory/Cardinal/Finite.lean @@ -5,7 +5,8 @@ Authors: Aaron Anderson -/ import Mathlib.Data.ULift import Mathlib.Data.ZMod.Defs -import Mathlib.SetTheory.Cardinal.PartENat +import Mathlib.SetTheory.Cardinal.ToNat +import Mathlib.SetTheory.Cardinal.ENat /-! # Finite Cardinality Functions @@ -321,103 +322,3 @@ theorem one_lt_card_iff_nontrivial (α : Type*) : 1 < card α ↔ Nontrivial α simp only [ENat.card, Nat.cast_one] end ENat - - -namespace PartENat - -/-- `PartENat.card α` is the cardinality of `α` as an extended natural number. - If `α` is infinite, `PartENat.card α = ⊤`. -/ -@[deprecated ENat.card (since := "2024-11-30")] -def card (α : Type*) : PartENat := - toPartENat (mk α) - --- The remainder of this section is about the deprecated `PartENat.card`. -set_option linter.deprecated false - -@[simp, deprecated ENat.card_eq_coe_fintype_card (since := "2024-11-30")] -theorem card_eq_coe_fintype_card [Fintype α] : card α = Fintype.card α := - mk_toPartENat_eq_coe_card - -@[simp, deprecated ENat.card_eq_top_of_infinite (since := "2024-11-30")] -theorem card_eq_top_of_infinite [Infinite α] : card α = ⊤ := - mk_toPartENat_of_infinite - -@[simp, deprecated ENat.card_sum (since := "2024-11-30")] -theorem card_sum (α β : Type*) : - PartENat.card (α ⊕ β) = PartENat.card α + PartENat.card β := by - simp only [PartENat.card, Cardinal.mk_sum, map_add, Cardinal.toPartENat_lift] - -@[deprecated ENat.card_congr (since := "2024-11-30")] -theorem card_congr {α : Type*} {β : Type*} (f : α ≃ β) : PartENat.card α = PartENat.card β := - Cardinal.toPartENat_congr f - -@[simp, deprecated ENat.card_ulift (since := "2024-11-30")] -lemma card_ulift (α : Type*) : card (ULift α) = card α := card_congr Equiv.ulift - -@[simp, deprecated ENat.card_plift (since := "2024-11-30")] -lemma card_plift (α : Type*) : card (PLift α) = card α := card_congr Equiv.plift - -@[deprecated ENat.card_image_of_injOn (since := "2024-11-30")] -theorem card_image_of_injOn {α : Type u} {β : Type v} {f : α → β} {s : Set α} (h : Set.InjOn f s) : - card (f '' s) = card s := - card_congr (Equiv.Set.imageOfInjOn f s h).symm - -@[deprecated ENat.card_image_of_injective (since := "2024-11-30")] -theorem card_image_of_injective {α : Type u} {β : Type v} (f : α → β) (s : Set α) - (h : Function.Injective f) : card (f '' s) = card s := card_image_of_injOn h.injOn - --- Should I keep the 6 following lemmas ? --- TODO: Add ofNat, zero, and one versions for simp confluence -@[simp, deprecated Cardinal.natCast_le_toENat_iff (since := "2024-11-30")] -theorem _root_.Cardinal.natCast_le_toPartENat_iff {n : ℕ} {c : Cardinal} : - ↑n ≤ toPartENat c ↔ ↑n ≤ c := by - rw [← toPartENat_natCast n, toPartENat_le_iff_of_le_aleph0 (le_of_lt (nat_lt_aleph0 n))] - -@[simp, deprecated Cardinal.toENat_le_natCast_iff (since := "2024-11-30")] -theorem _root_.Cardinal.toPartENat_le_natCast_iff {c : Cardinal} {n : ℕ} : - toPartENat c ≤ n ↔ c ≤ n := by - rw [← toPartENat_natCast n, toPartENat_le_iff_of_lt_aleph0 (nat_lt_aleph0 n)] - -@[simp, deprecated Cardinal.natCast_eq_toENat_iff (since := "2024-11-30")] -theorem _root_.Cardinal.natCast_eq_toPartENat_iff {n : ℕ} {c : Cardinal} : - ↑n = toPartENat c ↔ ↑n = c := by - rw [le_antisymm_iff, le_antisymm_iff, Cardinal.toPartENat_le_natCast_iff, - Cardinal.natCast_le_toPartENat_iff] - -@[simp, deprecated Cardinal.toENat_eq_natCast_iff (since := "2024-11-30")] -theorem _root_.Cardinal.toPartENat_eq_natCast_iff {c : Cardinal} {n : ℕ} : - Cardinal.toPartENat c = n ↔ c = n := by -rw [eq_comm, Cardinal.natCast_eq_toPartENat_iff, eq_comm] - -@[simp, deprecated Cardinal.natCast_lt_toENat_iff (since := "2024-11-30")] -theorem _root_.Cardinal.natCast_lt_toPartENat_iff {n : ℕ} {c : Cardinal} : - ↑n < toPartENat c ↔ ↑n < c := by - simp only [← not_le, Cardinal.toPartENat_le_natCast_iff] - -@[simp, deprecated Cardinal.toENat_lt_natCast_iff (since := "2024-11-30")] -theorem _root_.Cardinal.toPartENat_lt_natCast_iff {n : ℕ} {c : Cardinal} : - toPartENat c < ↑n ↔ c < ↑n := by - simp only [← not_le, Cardinal.natCast_le_toPartENat_iff] - -@[deprecated ENat.card_eq_zero_iff_empty (since := "2024-11-30")] -theorem card_eq_zero_iff_empty (α : Type*) : card α = 0 ↔ IsEmpty α := by - rw [← Cardinal.mk_eq_zero_iff] - conv_rhs => rw [← Nat.cast_zero] - simp only [← Cardinal.toPartENat_eq_natCast_iff] - simp only [PartENat.card, Nat.cast_zero] - -@[deprecated ENat.card_le_one_iff_subsingleton (since := "2024-11-30")] -theorem card_le_one_iff_subsingleton (α : Type*) : card α ≤ 1 ↔ Subsingleton α := by - rw [← le_one_iff_subsingleton] - conv_rhs => rw [← Nat.cast_one] - rw [← Cardinal.toPartENat_le_natCast_iff] - simp only [PartENat.card, Nat.cast_one] - -@[deprecated ENat.one_lt_card_iff_nontrivial (since := "2024-11-30")] -theorem one_lt_card_iff_nontrivial (α : Type*) : 1 < card α ↔ Nontrivial α := by - rw [← Cardinal.one_lt_iff_nontrivial] - conv_rhs => rw [← Nat.cast_one] - rw [← natCast_lt_toPartENat_iff] - simp only [PartENat.card, Nat.cast_one] - -end PartENat From 134c6ee3da5185da90b69d05697c85bfba57e82e Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Tue, 3 Dec 2024 06:18:14 +0000 Subject: [PATCH 484/829] chore: add script to merge lean testing PRs (#19531) --- scripts/README.md | 4 +++ scripts/merge-lean-testing-pr.sh | 44 ++++++++++++++++++++++++++++++++ 2 files changed, 48 insertions(+) create mode 100755 scripts/merge-lean-testing-pr.sh diff --git a/scripts/README.md b/scripts/README.md index 3962da0193667..8cc2baee843ad 100644 --- a/scripts/README.md +++ b/scripts/README.md @@ -67,6 +67,10 @@ to learn about it as well! - finally, merge the new branch back into `nightly-testing`, if conflict resolution was required. If there are merge conflicts, it pauses and asks for help from the human driver. +- `merge-lean-testing-pr.sh` takes a PR number `NNNN` as argument, + and attempts to merge the branch `lean-pr-testing-NNNN` into `master`. + It will resolve conflicts in `lean-toolchain`, `lakefile.lean`, and `lake-manifest.json`. + If there are more conflicts, it will bail. **Managing and tracking technical debt** - `technical-debt-metrics.sh` diff --git a/scripts/merge-lean-testing-pr.sh b/scripts/merge-lean-testing-pr.sh new file mode 100755 index 0000000000000..7a07419dbe47f --- /dev/null +++ b/scripts/merge-lean-testing-pr.sh @@ -0,0 +1,44 @@ +#!/usr/bin/env bash + +if [ "$#" -ne 1 ]; then + echo "Usage: $0 " + exit 1 +fi + +PR_NUMBER=$1 +BRANCH_NAME="lean-pr-testing-$PR_NUMBER" + +git checkout nightly-testing +git pull + +if ! git merge origin/$BRANCH_NAME; then + echo "Merge conflicts detected. Resolving conflicts in favor of current version..." + git checkout --ours lean-toolchain lakefile.lean lake-manifest.json + git add lean-toolchain lakefile.lean lake-manifest.json +fi + +sed -i "s/$BRANCH_NAME/nightly-testing/g" lakefile.lean +git add lakefile.lean + +# Check for merge conflicts +if git ls-files -u | grep -q '^'; then + echo "Merge conflicts detected. Please resolve conflicts manually." + git status + exit 1 +fi + +if ! lake update; then + echo "Lake update failed. Please resolve conflicts manually." + git status + exit 1 +fi + +# Attempt to commit. This will fail if there are conflicts. +if git commit -m "merge $BRANCH_NAME"; then + echo "Merge successful." + exit 0 +else + echo "Merge failed. Please resolve conflicts manually." + git status + exit 1 +fi From 5358a537cab9e63ab57d9ec41b64e112e275315d Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Tue, 3 Dec 2024 07:56:49 +0000 Subject: [PATCH 485/829] chore(Data/Int): move `Pi.instIntCast` (#19612) This instance was found in `Data.Int.Cast.Lemmas` and is used quite a few times without all the theory in those files. Co-authored-by: Anne Baanen --- Mathlib.lean | 1 + Mathlib/Data/Int/Cast/Lemmas.lean | 22 ---------- Mathlib/Data/Int/Cast/Pi.lean | 44 +++++++++++++++++++ Mathlib/Data/Int/CharZero.lean | 1 + Mathlib/Data/Matrix/Diagonal.lean | 2 +- .../MeasureTheory/MeasurableSpace/Basic.lean | 5 ++- Mathlib/Order/Filter/Germ/Basic.lean | 3 +- 7 files changed, 52 insertions(+), 26 deletions(-) create mode 100644 Mathlib/Data/Int/Cast/Pi.lean diff --git a/Mathlib.lean b/Mathlib.lean index 09bc8fed5e198..f1a2458fab371 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2526,6 +2526,7 @@ import Mathlib.Data.Int.Cast.Basic import Mathlib.Data.Int.Cast.Defs import Mathlib.Data.Int.Cast.Field import Mathlib.Data.Int.Cast.Lemmas +import Mathlib.Data.Int.Cast.Pi import Mathlib.Data.Int.Cast.Prod import Mathlib.Data.Int.CharZero import Mathlib.Data.Int.ConditionallyCompleteOrder diff --git a/Mathlib/Data/Int/Cast/Lemmas.lean b/Mathlib/Data/Int/Cast/Lemmas.lean index d887d52523d54..5610ebe757d4d 100644 --- a/Mathlib/Data/Int/Cast/Lemmas.lean +++ b/Mathlib/Data/Int/Cast/Lemmas.lean @@ -366,25 +366,3 @@ end NonAssocRing @[simp] theorem Int.castRingHom_int : Int.castRingHom ℤ = RingHom.id ℤ := (RingHom.id ℤ).eq_intCast'.symm - -namespace Pi - -variable {π : ι → Type*} [∀ i, IntCast (π i)] - -instance instIntCast : IntCast (∀ i, π i) where intCast n _ := n - -theorem intCast_apply (n : ℤ) (i : ι) : (n : ∀ i, π i) i = n := - rfl - -@[simp] -theorem intCast_def (n : ℤ) : (n : ∀ i, π i) = fun _ => ↑n := - rfl - -@[deprecated (since := "2024-04-05")] alias int_apply := intCast_apply -@[deprecated (since := "2024-04-05")] alias coe_int := intCast_def - -end Pi - -theorem Sum.elim_intCast_intCast {α β γ : Type*} [IntCast γ] (n : ℤ) : - Sum.elim (n : α → γ) (n : β → γ) = n := - Sum.elim_lam_const_lam_const (γ := γ) n diff --git a/Mathlib/Data/Int/Cast/Pi.lean b/Mathlib/Data/Int/Cast/Pi.lean new file mode 100644 index 0000000000000..cfdb7b2de7b54 --- /dev/null +++ b/Mathlib/Data/Int/Cast/Pi.lean @@ -0,0 +1,44 @@ +/- +Copyright (c) 2017 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro +-/ +import Batteries.Tactic.Alias +import Mathlib.Data.Int.Notation +import Mathlib.Tactic.TypeStar +import Mathlib.Util.AssertExists + +/-! +# Cast of integers to function types + +This file provides a (pointwise) cast from `ℤ` to function types. + +## Main declarations + +* `Pi.instIntCast`: map `n : ℤ` to the constant function `n : ∀ i, π i` +-/ + +assert_not_exists OrderedCommMonoid +assert_not_exists RingHom + +namespace Pi + +variable {ι : Type*} {π : ι → Type*} [∀ i, IntCast (π i)] + +instance instIntCast : IntCast (∀ i, π i) where intCast n _ := n + +theorem intCast_apply (n : ℤ) (i : ι) : (n : ∀ i, π i) i = n := + rfl + +@[simp] +theorem intCast_def (n : ℤ) : (n : ∀ i, π i) = fun _ => ↑n := + rfl + +@[deprecated (since := "2024-04-05")] alias int_apply := intCast_apply +@[deprecated (since := "2024-04-05")] alias coe_int := intCast_def + +end Pi + +theorem Sum.elim_intCast_intCast {α β γ : Type*} [IntCast γ] (n : ℤ) : + Sum.elim (n : α → γ) (n : β → γ) = n := + Sum.elim_lam_const_lam_const (γ := γ) n diff --git a/Mathlib/Data/Int/CharZero.lean b/Mathlib/Data/Int/CharZero.lean index 11a1e6a858d9a..49a66742d5cd6 100644 --- a/Mathlib/Data/Int/CharZero.lean +++ b/Mathlib/Data/Int/CharZero.lean @@ -6,6 +6,7 @@ Authors: Mario Carneiro import Mathlib.Algebra.Group.Support import Mathlib.Data.Int.Cast.Field import Mathlib.Data.Int.Cast.Lemmas +import Mathlib.Data.Int.Cast.Pi /-! # Injectivity of `Int.Cast` into characteristic zero rings and fields. diff --git a/Mathlib/Data/Matrix/Diagonal.lean b/Mathlib/Data/Matrix/Diagonal.lean index 6e226513da340..a1ae509a8b378 100644 --- a/Mathlib/Data/Matrix/Diagonal.lean +++ b/Mathlib/Data/Matrix/Diagonal.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Ellen Arlt. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Ellen Arlt, Blair Shi, Sean Leather, Mario Carneiro, Johan Commelin, Lu-Ming Zhang -/ -import Mathlib.Data.Int.Cast.Lemmas +import Mathlib.Data.Int.Cast.Pi import Mathlib.Data.Matrix.Defs import Mathlib.Data.Nat.Cast.Basic diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index 678a6983571d0..9693de1e0a8bf 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -4,14 +4,15 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ import Mathlib.Data.Finset.Update -import Mathlib.Data.Int.Cast.Lemmas +import Mathlib.Data.Int.Cast.Pi +import Mathlib.Data.Nat.Cast.Basic import Mathlib.Data.Prod.TProd import Mathlib.Data.Set.UnionLift import Mathlib.GroupTheory.Coset.Defs import Mathlib.MeasureTheory.MeasurableSpace.Instances import Mathlib.Order.Filter.SmallSets -import Mathlib.Tactic.FinCases import Mathlib.Order.LiminfLimsup +import Mathlib.Tactic.FinCases /-! # Measurable spaces and measurable functions diff --git a/Mathlib/Order/Filter/Germ/Basic.lean b/Mathlib/Order/Filter/Germ/Basic.lean index c9ad7754c4946..b7df9bc2a0c66 100644 --- a/Mathlib/Order/Filter/Germ/Basic.lean +++ b/Mathlib/Order/Filter/Germ/Basic.lean @@ -5,7 +5,8 @@ Authors: Yury Kudryashov, Abhimanyu Pallavi Sudhir -/ import Mathlib.Algebra.Module.Pi import Mathlib.Algebra.Order.Monoid.Unbundled.ExistsOfLE -import Mathlib.Data.Int.Cast.Lemmas +import Mathlib.Data.Int.Cast.Pi +import Mathlib.Data.Nat.Cast.Basic import Mathlib.Order.Filter.Tendsto /-! From 1443a7f471cfdde855509db68de15981c2d73096 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Tue, 3 Dec 2024 08:47:28 +0000 Subject: [PATCH 486/829] chore(Mathlib/Algebra/MvPolynomial/Degrees): rename `degreeOf_mul_X_eq` (#19654) Co-authored-by: Moritz Firsching --- Mathlib/Algebra/MvPolynomial/Degrees.lean | 14 +++++++++----- 1 file changed, 9 insertions(+), 5 deletions(-) diff --git a/Mathlib/Algebra/MvPolynomial/Degrees.lean b/Mathlib/Algebra/MvPolynomial/Degrees.lean index e0e9da35be358..45521d9bd0bb7 100644 --- a/Mathlib/Algebra/MvPolynomial/Degrees.lean +++ b/Mathlib/Algebra/MvPolynomial/Degrees.lean @@ -282,28 +282,32 @@ theorem degreeOf_pow_le (i : σ) (p : MvPolynomial σ R) (n : ℕ) : degreeOf i (p ^ n) ≤ n * degreeOf i p := by simpa using degreeOf_prod_le i (Finset.range n) (fun _ => p) -theorem degreeOf_mul_X_ne {i j : σ} (f : MvPolynomial σ R) (h : i ≠ j) : +theorem degreeOf_mul_X_of_ne {i j : σ} (f : MvPolynomial σ R) (h : i ≠ j) : degreeOf i (f * X j) = degreeOf i f := by classical simp only [degreeOf_eq_sup i, support_mul_X, Finset.sup_map] congr ext - simp only [Finsupp.single, Nat.one_ne_zero, add_right_eq_self, addRightEmbedding_apply, coe_mk, + simp only [Finsupp.single, add_right_eq_self, addRightEmbedding_apply, coe_mk, Pi.add_apply, comp_apply, ite_eq_right_iff, Finsupp.coe_add, Pi.single_eq_of_ne h] -theorem degreeOf_mul_X_eq (j : σ) (f : MvPolynomial σ R) : +@[deprecated (since := "2024-12-01")] alias degreeOf_mul_X_ne := degreeOf_mul_X_of_ne + +theorem degreeOf_mul_X_self (j : σ) (f : MvPolynomial σ R) : degreeOf j (f * X j) ≤ degreeOf j f + 1 := by classical simp only [degreeOf] apply (Multiset.count_le_of_le j (degrees_mul f (X j))).trans simp only [Multiset.count_add, add_le_add_iff_left] - convert Multiset.count_le_of_le j (degrees_X' (R := R) j) + convert Multiset.count_le_of_le j <| degrees_X' j rw [Multiset.count_singleton_self] +@[deprecated (since := "2024-12-01")] alias degreeOf_mul_X_eq := degreeOf_mul_X_self + theorem degreeOf_mul_X_eq_degreeOf_add_one_iff (j : σ) (f : MvPolynomial σ R) : degreeOf j (f * X j) = degreeOf j f + 1 ↔ f ≠ 0 := by refine ⟨fun h => by by_contra ha; simp [ha] at h, fun h => ?_⟩ - apply Nat.le_antisymm (degreeOf_mul_X_eq j f) + apply Nat.le_antisymm (degreeOf_mul_X_self j f) have : (f.support.sup fun m ↦ m j) + 1 = (f.support.sup fun m ↦ (m j + 1)) := Finset.comp_sup_eq_sup_comp_of_nonempty @Nat.succ_le_succ (support_nonempty.mpr h) simp only [degreeOf_eq_sup, support_mul_X, this] From deb1faaee197406368b6bc1f680c5c242c955045 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 3 Dec 2024 10:28:54 +0000 Subject: [PATCH 487/829] =?UTF-8?q?feat(Pointwise):=20`closure=20(s=20^=20?= =?UTF-8?q?n)=20=3D=20closure=20s`=20if=20`1=20=E2=88=88=20s`=20(#19577)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From GrowthInGroups (LeanCamCombi) --- Mathlib/Algebra/Group/Subgroup/Pointwise.lean | 15 +++++++++++++++ Mathlib/Algebra/Group/Submonoid/Pointwise.lean | 15 +++++++++++++++ 2 files changed, 30 insertions(+) diff --git a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean index d5827d8bec2d1..b5e3e4461137b 100644 --- a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean +++ b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean @@ -202,6 +202,21 @@ theorem closure_mul_le (S T : Set G) : closure (S * T) ≤ closure S ⊔ closure (closure S ⊔ closure T).mul_mem (SetLike.le_def.mp le_sup_left <| subset_closure hs) (SetLike.le_def.mp le_sup_right <| subset_closure ht) +@[to_additive] +lemma closure_pow_le : ∀ {n}, n ≠ 0 → closure (s ^ n) ≤ closure s + | 1, _ => by simp + | n + 2, _ => + calc + closure (s ^ (n + 2)) + _ = closure (s ^ (n + 1) * s) := by rw [pow_succ] + _ ≤ closure (s ^ (n + 1)) ⊔ closure s := closure_mul_le .. + _ ≤ closure s ⊔ closure s := by gcongr ?_ ⊔ _; exact closure_pow_le n.succ_ne_zero + _ = closure s := sup_idem _ + +@[to_additive] +lemma closure_pow {n : ℕ} (hs : 1 ∈ s) (hn : n ≠ 0) : closure (s ^ n) = closure s := + (closure_pow_le hn).antisymm <| by gcongr; exact subset_pow hs hn + @[to_additive] theorem sup_eq_closure_mul (H K : Subgroup G) : H ⊔ K = closure ((H : Set G) * (K : Set G)) := le_antisymm diff --git a/Mathlib/Algebra/Group/Submonoid/Pointwise.lean b/Mathlib/Algebra/Group/Submonoid/Pointwise.lean index cc8c81e07b75a..596c1f0f1ca02 100644 --- a/Mathlib/Algebra/Group/Submonoid/Pointwise.lean +++ b/Mathlib/Algebra/Group/Submonoid/Pointwise.lean @@ -96,6 +96,21 @@ theorem closure_mul_le (S T : Set M) : closure (S * T) ≤ closure S ⊔ closure (closure S ⊔ closure T).mul_mem (SetLike.le_def.mp le_sup_left <| subset_closure hs) (SetLike.le_def.mp le_sup_right <| subset_closure ht) +@[to_additive] +lemma closure_pow_le : ∀ {n}, n ≠ 0 → closure (s ^ n) ≤ closure s + | 1, _ => by simp + | n + 2, _ => + calc + closure (s ^ (n + 2)) + _ = closure (s ^ (n + 1) * s) := by rw [pow_succ] + _ ≤ closure (s ^ (n + 1)) ⊔ closure s := closure_mul_le .. + _ ≤ closure s ⊔ closure s := by gcongr ?_ ⊔ _; exact closure_pow_le n.succ_ne_zero + _ = closure s := sup_idem _ + +@[to_additive] +lemma closure_pow {n : ℕ} (hs : 1 ∈ s) (hn : n ≠ 0) : closure (s ^ n) = closure s := + (closure_pow_le hn).antisymm <| by gcongr; exact subset_pow hs hn + @[to_additive] theorem sup_eq_closure_mul (H K : Submonoid M) : H ⊔ K = closure ((H : Set M) * (K : Set M)) := le_antisymm From 0ddc4d7ff033404e81f865e12afac01b5bdf0925 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Tue, 3 Dec 2024 10:57:46 +0000 Subject: [PATCH 488/829] feat: prove Dini's theorem for uniform convergence (#19068) --- Mathlib.lean | 1 + Mathlib/Topology/Algebra/Group/Basic.lean | 6 + Mathlib/Topology/UniformSpace/Basic.lean | 4 + Mathlib/Topology/UniformSpace/Dini.lean | 159 ++++++++++++++++++++++ 4 files changed, 170 insertions(+) create mode 100644 Mathlib/Topology/UniformSpace/Dini.lean diff --git a/Mathlib.lean b/Mathlib.lean index f1a2458fab371..42c0594b705cc 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -5353,6 +5353,7 @@ import Mathlib.Topology.UniformSpace.CompactConvergence import Mathlib.Topology.UniformSpace.CompareReals import Mathlib.Topology.UniformSpace.CompleteSeparated import Mathlib.Topology.UniformSpace.Completion +import Mathlib.Topology.UniformSpace.Dini import Mathlib.Topology.UniformSpace.Equicontinuity import Mathlib.Topology.UniformSpace.Equiv import Mathlib.Topology.UniformSpace.HeineCantor diff --git a/Mathlib/Topology/Algebra/Group/Basic.lean b/Mathlib/Topology/Algebra/Group/Basic.lean index 4b003dcf49f11..6746a7d253f9f 100644 --- a/Mathlib/Topology/Algebra/Group/Basic.lean +++ b/Mathlib/Topology/Algebra/Group/Basic.lean @@ -233,6 +233,9 @@ theorem ContinuousWithinAt.inv (hf : ContinuousWithinAt f s x) : ContinuousWithinAt (fun x => (f x)⁻¹) s x := Filter.Tendsto.inv hf +@[to_additive] +instance OrderDual.instContinuousInv : ContinuousInv Gᵒᵈ := ‹ContinuousInv G› + @[to_additive] instance Prod.continuousInv [TopologicalSpace H] [Inv H] [ContinuousInv H] : ContinuousInv (G × H) := @@ -536,6 +539,9 @@ end OrderedCommGroup instance [TopologicalSpace H] [Group H] [TopologicalGroup H] : TopologicalGroup (G × H) where continuous_inv := continuous_inv.prodMap continuous_inv +@[to_additive] +instance OrderDual.instTopologicalGroup : TopologicalGroup Gᵒᵈ where + @[to_additive] instance Pi.topologicalGroup {C : β → Type*} [∀ b, TopologicalSpace (C b)] [∀ b, Group (C b)] [∀ b, TopologicalGroup (C b)] : TopologicalGroup (∀ b, C b) where diff --git a/Mathlib/Topology/UniformSpace/Basic.lean b/Mathlib/Topology/UniformSpace/Basic.lean index 9f28773dd89af..ef9363d4616d6 100644 --- a/Mathlib/Topology/UniformSpace/Basic.lean +++ b/Mathlib/Topology/UniformSpace/Basic.lean @@ -1217,6 +1217,10 @@ theorem UniformContinuous.continuous [UniformSpace α] [UniformSpace β] {f : α instance ULift.uniformSpace [UniformSpace α] : UniformSpace (ULift α) := UniformSpace.comap ULift.down ‹_› +/-- Uniform space structure on `αᵒᵈ`. -/ +instance OrderDual.instUniformSpace [UniformSpace α] : UniformSpace (αᵒᵈ) := + ‹UniformSpace α› + section UniformContinuousInfi -- Porting note: renamed for dot notation; add an `iff` lemma? diff --git a/Mathlib/Topology/UniformSpace/Dini.lean b/Mathlib/Topology/UniformSpace/Dini.lean new file mode 100644 index 0000000000000..9e90b66608f20 --- /dev/null +++ b/Mathlib/Topology/UniformSpace/Dini.lean @@ -0,0 +1,159 @@ +/- +Copyright (c) 2024 Jireh Loreaux. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jireh Loreaux +-/ +import Mathlib.Analysis.Normed.Order.Lattice +import Mathlib.Topology.ContinuousMap.Ordered +import Mathlib.Topology.UniformSpace.CompactConvergence + +/-! # Dini's Theorem + +This file proves Dini's theorem, which states that if `F n` is a monotone increasing sequence of +continuous real-valued functions on a compact set `s` converging pointwise to a continuous function +`f`, then `F n` converges uniformly to `f`. + +We generalize the codomain from `ℝ` to a normed lattice additive commutative group `G`. +This theorem is true in a different generality as well: when `G` is a linearly ordered topological +group with the order topology. This weakens the norm assumption, in exchange for strengthening to +a linear order. This separate generality is not included in this file, but that generality was +included in initial drafts of the original +[PR #19068](https://github.com/leanprover-community/mathlib4/pull/19068) and can be recovered if +necessary. + +The key idea of the proof is to use a particular basis of `𝓝 0` which consists of open sets that +are somehow monotone in the sense that if `s` is in the basis, and `0 ≤ x ≤ y`, then +`y ∈ s → x ∈ s`, and so the proof would work on any topological ordered group possessing +such a basis. In the case of a linearly ordered topological group with the order topology, this +basis is `nhds_basis_Ioo`. In the case of a normed lattice additive commutative group, this basis +is `nhds_basis_ball`, and the fact that this basis satisfies the monotonicity criterion +corresponds to `HasSolidNorm`. +-/ + +open Filter Topology + +variable {ι α G : Type*} [Preorder ι] [TopologicalSpace α] [NormedLatticeAddCommGroup G] + +section Unbundled + +open Metric + +variable {F : ι → α → G} {f : α → G} + +namespace Monotone + +/-- **Dini's theorem** If `F n` is a monotone increasing collection of continuous functions +converging pointwise to a continuous function `f`, then `F n` converges locally uniformly to `f`. -/ +lemma tendstoLocallyUniformly_of_forall_tendsto + (hF_cont : ∀ i, Continuous (F i)) (hF_mono : Monotone F) (hf : Continuous f) + (h_tendsto : ∀ x, Tendsto (F · x) atTop (𝓝 (f x))) : + TendstoLocallyUniformly F f atTop := by + refine (atTop : Filter ι).eq_or_neBot.elim (fun h ↦ ?eq_bot) (fun _ ↦ ?_) + case eq_bot => simp [h, tendstoLocallyUniformly_iff_forall_tendsto] + have F_le_f (x : α) (n : ι) : F n x ≤ f x := by + refine ge_of_tendsto (h_tendsto x) ?_ + filter_upwards [Ici_mem_atTop n] with m hnm + exact hF_mono hnm x + simp_rw [Metric.tendstoLocallyUniformly_iff, dist_eq_norm'] + intro ε ε_pos x + simp_rw +singlePass [tendsto_iff_norm_sub_tendsto_zero] at h_tendsto + obtain ⟨n, hn⟩ := (h_tendsto x).eventually (eventually_lt_nhds ε_pos) |>.exists + refine ⟨{y | ‖F n y - f y‖ < ε}, ⟨isOpen_lt (by fun_prop) continuous_const |>.mem_nhds hn, ?_⟩⟩ + filter_upwards [eventually_ge_atTop n] with m hnm z hz + refine norm_le_norm_of_abs_le_abs ?_ |>.trans_lt hz + simp only [abs_of_nonpos (sub_nonpos_of_le (F_le_f _ _)), neg_sub, sub_le_sub_iff_left] + exact hF_mono hnm z + +/-- **Dini's theorem** If `F n` is a monotone increasing collection of continuous functions on a +set `s` converging pointwise to a continuous funciton `f`, then `F n` converges locally uniformly +to `f`. -/ +lemma tendstoLocallyUniformlyOn_of_forall_tendsto {s : Set α} + (hF_cont : ∀ i, ContinuousOn (F i) s) (hF_mono : ∀ x ∈ s, Monotone (F · x)) + (hf : ContinuousOn f s) (h_tendsto : ∀ x ∈ s, Tendsto (F · x) atTop (𝓝 (f x))) : + TendstoLocallyUniformlyOn F f atTop s := by + rw [tendstoLocallyUniformlyOn_iff_tendstoLocallyUniformly_comp_coe] + exact tendstoLocallyUniformly_of_forall_tendsto (hF_cont · |>.restrict) + (fun _ _ h x ↦ hF_mono _ x.2 h) hf.restrict (fun x ↦ h_tendsto x x.2) + +/-- **Dini's theorem** If `F n` is a monotone increasing collection of continuous functions on a +compact space converging pointwise to a continuous function `f`, then `F n` converges uniformly to +`f`. -/ +lemma tendstoUniformly_of_forall_tendsto [CompactSpace α] (hF_cont : ∀ i, Continuous (F i)) + (hF_mono : Monotone F) (hf : Continuous f) (h_tendsto : ∀ x, Tendsto (F · x) atTop (𝓝 (f x))) : + TendstoUniformly F f atTop := + tendstoLocallyUniformly_iff_tendstoUniformly_of_compactSpace.mp <| + tendstoLocallyUniformly_of_forall_tendsto hF_cont hF_mono hf h_tendsto + +/-- **Dini's theorem** If `F n` is a monotone increasing collection of continuous functions on a +compact set `s` converging pointwise to a continuous function `f`, then `F n` converges uniformly +to `f`. -/ +lemma tendstoUniformlyOn_of_forall_tendsto {s : Set α} (hs : IsCompact s) + (hF_cont : ∀ i, ContinuousOn (F i) s) (hF_mono : ∀ x ∈ s, Monotone (F · x)) + (hf : ContinuousOn f s) (h_tendsto : ∀ x ∈ s, Tendsto (F · x) atTop (𝓝 (f x))) : + TendstoUniformlyOn F f atTop s := + tendstoLocallyUniformlyOn_iff_tendstoUniformlyOn_of_compact hs |>.mp <| + tendstoLocallyUniformlyOn_of_forall_tendsto hF_cont hF_mono hf h_tendsto + +end Monotone + +namespace Antitone + +/-- **Dini's theorem** If `F n` is a monotone decreasing collection of continuous functions on a +converging pointwise to a continuous function `f`, then `F n` converges locally uniformly to `f`. -/ +lemma tendstoLocallyUniformly_of_forall_tendsto + (hF_cont : ∀ i, Continuous (F i)) (hF_anti : Antitone F) (hf : Continuous f) + (h_tendsto : ∀ x, Tendsto (F · x) atTop (𝓝 (f x))) : + TendstoLocallyUniformly F f atTop := + Monotone.tendstoLocallyUniformly_of_forall_tendsto (G := Gᵒᵈ) hF_cont hF_anti hf h_tendsto + +/-- **Dini's theorem** If `F n` is a monotone decreasing collection of continuous functions on a +set `s` converging pointwise to a continuous function `f`, then `F n` converges locally uniformly +to `f`. -/ +lemma tendstoLocallyUniformlyOn_of_forall_tendsto {s : Set α} + (hF_cont : ∀ i, ContinuousOn (F i) s) (hF_anti : ∀ x ∈ s, Antitone (F · x)) + (hf : ContinuousOn f s) (h_tendsto : ∀ x ∈ s, Tendsto (F · x) atTop (𝓝 (f x))) : + TendstoLocallyUniformlyOn F f atTop s := + Monotone.tendstoLocallyUniformlyOn_of_forall_tendsto (G := Gᵒᵈ) hF_cont hF_anti hf h_tendsto + +/-- **Dini's theorem** If `F n` is a monotone decreasing collection of continuous functions on a +compact space converging pointwise to a continuous function `f`, then `F n` converges uniformly +to `f`. -/ +lemma tendstoUniformly_of_forall_tendsto [CompactSpace α] (hF_cont : ∀ i, Continuous (F i)) + (hF_anti : Antitone F) (hf : Continuous f) (h_tendsto : ∀ x, Tendsto (F · x) atTop (𝓝 (f x))) : + TendstoUniformly F f atTop := + Monotone.tendstoUniformly_of_forall_tendsto (G := Gᵒᵈ) hF_cont hF_anti hf h_tendsto + +/-- **Dini's theorem** If `F n` is a monotone decreasing collection of continuous functions on a +compact set `s` converging pointwise to a continuous `f`, then `F n` converges uniformly to `f`. -/ +lemma tendstoUniformlyOn_of_forall_tendsto {s : Set α} (hs : IsCompact s) + (hF_cont : ∀ i, ContinuousOn (F i) s) (hF_anti : ∀ x ∈ s, Antitone (F · x)) + (hf : ContinuousOn f s) (h_tendsto : ∀ x ∈ s, Tendsto (F · x) atTop (𝓝 (f x))) : + TendstoUniformlyOn F f atTop s := + Monotone.tendstoUniformlyOn_of_forall_tendsto (G := Gᵒᵈ) hs hF_cont hF_anti hf h_tendsto + +end Antitone + +end Unbundled + +namespace ContinuousMap + +variable {F : ι → C(α, G)} {f : C(α, G)} + +/-- **Dini's theorem** If `F n` is a monotone increasing collection of continuous functions +converging pointwise to a continuous function `f`, then `F n` converges to `f` in the +compact-open topology. -/ +lemma tendsto_of_monotone_of_pointwise (hF_mono : Monotone F) + (h_tendsto : ∀ x, Tendsto (F · x) atTop (𝓝 (f x))) : + Tendsto F atTop (𝓝 f) := + tendsto_of_tendstoLocallyUniformly <| + hF_mono.tendstoLocallyUniformly_of_forall_tendsto (F · |>.continuous) f.continuous h_tendsto + +/-- **Dini's theorem** If `F n` is a monotone decreasing collection of continuous functions +converging pointwise to a continuous function `f`, then `F n` converges to `f` in the +compact-open topology. -/ +lemma tendsto_of_antitone_of_pointwise (hF_anti : Antitone F) + (h_tendsto : ∀ x, Tendsto (F · x) atTop (𝓝 (f x))) : + Tendsto F atTop (𝓝 f) := + tendsto_of_monotone_of_pointwise (G := Gᵒᵈ) hF_anti h_tendsto + +end ContinuousMap From 19e3a19ba434faf8d3559b7acae30c5b064704fa Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 3 Dec 2024 10:57:48 +0000 Subject: [PATCH 489/829] chore(Imo/Imo1964Q1): minor clean-up (#19691) This PR does the following: - remove some trivial sectioning - explain the solution better - remove `ProblemPredicate`, which was only used once - replace `ring` by a short `rw` (to minimize imports) --- Archive/Imo/Imo1964Q1.lean | 30 +++++++----------------------- 1 file changed, 7 insertions(+), 23 deletions(-) diff --git a/Archive/Imo/Imo1964Q1.lean b/Archive/Imo/Imo1964Q1.lean index dda0e7f1e11cb..6a3e8cba29dfd 100644 --- a/Archive/Imo/Imo1964Q1.lean +++ b/Archive/Imo/Imo1964Q1.lean @@ -5,25 +5,17 @@ Authors: Kevin Buzzard -/ import Mathlib.Tactic.IntervalCases import Mathlib.Data.Nat.ModEq -import Mathlib.Tactic.Ring /-! # IMO 1964 Q1 (a) Find all positive integers $n$ for which $2^n-1$ is divisible by $7$. - (b) Prove that there is no positive integer $n$ for which $2^n+1$ is divisible by $7$. -We define a predicate for the solutions in (a), and prove that it is the set of positive -integers which are a multiple of 3. --/ - - -/-! -## Intermediate lemmas +For (a), we find that the order of $2$ mod $7$ is $3$. Therefore for (b), it suffices to check +$n = 0, 1, 2$. -/ - open Nat namespace Imo1964Q1 @@ -33,17 +25,13 @@ theorem two_pow_mod_seven (n : ℕ) : 2 ^ n ≡ 2 ^ (n % 3) [MOD 7] := calc 2 ^ n = 2 ^ (3 * (n / 3) + t) := by rw [Nat.div_add_mod] _ = (2 ^ 3) ^ (n / 3) * 2 ^ t := by rw [pow_add, pow_mul] _ ≡ 1 ^ (n / 3) * 2 ^ t [MOD 7] := by gcongr; decide - _ = 2 ^ t := by ring - -/-! -## The question --/ + _ = 2 ^ t := by rw [one_pow, one_mul] +end Imo1964Q1 -def ProblemPredicate (n : ℕ) : Prop := - 7 ∣ 2 ^ n - 1 +open Imo1964Q1 -theorem imo1964_q1a (n : ℕ) (_ : 0 < n) : ProblemPredicate n ↔ 3 ∣ n := by +theorem imo1964_q1a (n : ℕ) (_ : 0 < n) : 7 ∣ 2 ^ n - 1 ↔ 3 ∣ n := by let t := n % 3 have : t < 3 := Nat.mod_lt _ (by decide) calc 7 ∣ 2 ^ n - 1 ↔ 2 ^ n ≡ 1 [MOD 7] := by @@ -53,15 +41,11 @@ theorem imo1964_q1a (n : ℕ) (_ : 0 < n) : ProblemPredicate n ↔ 3 ∣ n := by _ ↔ t = 0 := by interval_cases t <;> decide _ ↔ 3 ∣ n := by rw [dvd_iff_mod_eq_zero] -end Imo1964Q1 - -open Imo1964Q1 - theorem imo1964_q1b (n : ℕ) : ¬7 ∣ 2 ^ n + 1 := by intro h let t := n % 3 have : t < 3 := Nat.mod_lt _ (by decide) have H : 2 ^ t + 1 ≡ 0 [MOD 7] := calc - 2 ^ t + 1 ≡ 2 ^ n + 1 [MOD 7 ] := by gcongr ?_ + 1; exact (two_pow_mod_seven n).symm + 2 ^ t + 1 ≡ 2 ^ n + 1 [MOD 7] := by gcongr ?_ + 1; exact (two_pow_mod_seven n).symm _ ≡ 0 [MOD 7] := h.modEq_zero_nat interval_cases t <;> contradiction From 1ec33a404152f54bbabe1162c3dfa356a6f8b0ab Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 3 Dec 2024 10:57:49 +0000 Subject: [PATCH 490/829] chore: cleanup set_option linter.deprecated (#19701) Remove several uses of deprecated lemmas in non-deprecated lemmas. Also convert all `set_option linter.deprecated false` to `set_option linter.deprecated false in`, so it is easier for us to count via regex what remains. I think ``` set_option linter\.deprecated false.*(?:(?:\n/--.*?-/\n(?!.*deprecated))|\n(?!.*deprecated|/--)) ``` is the correct regex to use now: it accounts properly for intervening doc-strings. It now reports that there are only ~~three~~ one! file~~s~~ that use `set_option linter.deprecated false` on non-deprecated declarations, namely: * ~~Mathlib.Data.List.Permutation~~ (fixed by moving some List theorems earlier) * ~~Mathlib.LinearAlgebra.CliffordAlgebra.Grading~~ (replace `AlgHom.map_zero` with just `map_zero`; only costs 2000 heartbeats) * Mathlib.SetTheory.Ordinal.Arithmetic (which is a mess) --- Mathlib/Algebra/BigOperators/Group/List.lean | 13 +---- Mathlib/Analysis/Analytic/Composition.lean | 9 ++-- Mathlib/Data/List/Basic.lean | 6 +-- Mathlib/Data/List/Lemmas.lean | 12 +++++ Mathlib/Data/List/Permutation.lean | 12 ++--- .../CliffordAlgebra/Grading.lean | 3 +- Mathlib/LinearAlgebra/Prod.lean | 16 +++++- Mathlib/Order/Extension/Well.lean | 5 +- Mathlib/Order/OmegaCompletePartialOrder.lean | 38 +++++++++----- Mathlib/Order/WellFounded.lean | 5 +- Mathlib/SetTheory/Cardinal/Aleph.lean | 25 ++++++++- Mathlib/SetTheory/Cardinal/Basic.lean | 11 ++-- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 52 +++++++++++++++++-- Mathlib/SetTheory/Ordinal/FixedPoint.lean | 24 ++++++++- Mathlib/SetTheory/Ordinal/Rank.lean | 6 ++- Mathlib/SetTheory/ZFC/Basic.lean | 21 ++++++-- Mathlib/Topology/Order/Basic.lean | 3 +- 17 files changed, 194 insertions(+), 67 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index 71307df2dbf18..0eb24a9825d47 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl, Floris van Doorn, Sébastien Gouëzel, Alex J. Best -/ import Mathlib.Algebra.Divisibility.Basic import Mathlib.Algebra.Group.Int +import Mathlib.Data.List.Lemmas import Mathlib.Data.List.Dedup import Mathlib.Data.List.Flatten import Mathlib.Data.List.Pairwise @@ -661,22 +662,10 @@ lemma mem_mem_ranges_iff_lt_sum (l : List ℕ) {n : ℕ} : (∃ s ∈ l.ranges, n ∈ s) ↔ n < l.sum := by rw [← mem_range, ← ranges_flatten, mem_flatten] -@[simp] -theorem length_flatMap (l : List α) (f : α → List β) : - length (List.flatMap l f) = sum (map (length ∘ f) l) := by - rw [List.flatMap, length_flatten, map_map] - @[deprecated (since := "2024-10-16")] alias length_bind := length_flatMap -lemma countP_flatMap (p : β → Bool) (l : List α) (f : α → List β) : - countP p (l.flatMap f) = sum (map (countP p ∘ f) l) := by - rw [List.flatMap, countP_flatten, map_map] - @[deprecated (since := "2024-10-16")] alias countP_bind := countP_flatMap -lemma count_flatMap [BEq β] (l : List α) (f : α → List β) (x : β) : - count x (l.flatMap f) = sum (map (count x ∘ f) l) := countP_flatMap _ _ _ - @[deprecated (since := "2024-10-16")] alias count_bind := count_flatMap /-- In a flatten, taking the first elements up to an index which is the sum of the lengths of the diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index 472d06dd79024..29001e428114c 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -1051,16 +1051,15 @@ theorem length_sigmaCompositionAux (a : Composition n) (b : Composition a.length show List.length ((splitWrtComposition a.blocks b)[i.1]) = blocksFun b i by rw [getElem_map_rev List.length, getElem_of_eq (map_length_splitWrtComposition _ _)]; rfl -set_option linter.deprecated false in theorem blocksFun_sigmaCompositionAux (a : Composition n) (b : Composition a.length) (i : Fin b.length) (j : Fin (blocksFun b i)) : blocksFun (sigmaCompositionAux a b ⟨i, (length_gather a b).symm ▸ i.2⟩) ⟨j, (length_sigmaCompositionAux a b i).symm ▸ j.2⟩ = - blocksFun a (embedding b i j) := - show get (get _ ⟨_, _⟩) ⟨_, _⟩ = a.blocks.get ⟨_, _⟩ by - rw [get_of_eq (get_splitWrtComposition _ _ _), get_drop', get_take']; rfl + blocksFun a (embedding b i j) := by + unfold sigmaCompositionAux + rw [blocksFun, get_eq_getElem, getElem_of_eq (getElem_splitWrtComposition _ _ _ _), + getElem_drop, getElem_take]; rfl -set_option linter.deprecated false in /-- Auxiliary lemma to prove that the composition of formal multilinear series is associative. Consider a composition `a` of `n` and a composition `b` of `a.length`. Grouping together some diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 5b5ca26520e79..6b8c7c19967e6 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -831,17 +831,13 @@ theorem get_reverse (l : List α) (i : Nat) (h1 h2) : dsimp omega -set_option linter.deprecated false - theorem get_reverse' (l : List α) (n) (hn') : l.reverse.get n = l.get ⟨l.length - 1 - n, hn'⟩ := by - rw [eq_comm] - convert get_reverse l.reverse n (by simpa) n.2 using 1 simp theorem eq_cons_of_length_one {l : List α} (h : l.length = 1) : l = [l.get ⟨0, by omega⟩] := by refine ext_get (by convert h) fun n h₁ h₂ => ?_ - simp only [get_singleton] + simp congr omega diff --git a/Mathlib/Data/List/Lemmas.lean b/Mathlib/Data/List/Lemmas.lean index 50f55ed3f3832..a6c98a38f2db9 100644 --- a/Mathlib/Data/List/Lemmas.lean +++ b/Mathlib/Data/List/Lemmas.lean @@ -17,6 +17,18 @@ variable {α β γ : Type*} namespace List +@[simp] +theorem length_flatMap (l : List α) (f : α → List β) : + length (List.flatMap l f) = sum (map (length ∘ f) l) := by + rw [List.flatMap, length_flatten, map_map] + +lemma countP_flatMap (p : β → Bool) (l : List α) (f : α → List β) : + countP p (l.flatMap f) = sum (map (countP p ∘ f) l) := by + rw [List.flatMap, countP_flatten, map_map] + +lemma count_flatMap [BEq β] (l : List α) (f : α → List β) (x : β) : + count x (l.flatMap f) = sum (map (count x ∘ f) l) := countP_flatMap _ _ _ + @[deprecated (since := "2024-08-20")] alias getElem_reverse' := getElem_reverse theorem tail_reverse_eq_reverse_dropLast (l : List α) : diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index f436730690267..49f6dd007cd72 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -3,7 +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.Data.List.Flatten +import Mathlib.Data.List.Lemmas import Mathlib.Data.Nat.Factorial.Basic import Mathlib.Data.List.Count import Mathlib.Data.List.Duplicate @@ -181,20 +181,18 @@ theorem mem_foldr_permutationsAux2 {t : α} {ts : List α} {r L : List (List α) simp only [mem_permutationsAux2', ← this, or_comm, and_left_comm, mem_append, mem_flatMap, append_assoc, cons_append, exists_prop] -set_option linter.deprecated false in theorem length_foldr_permutationsAux2 (t : α) (ts : List α) (r L : List (List α)) : length (foldr (fun y r => (permutationsAux2 t ts r y id).2) r L) = - Nat.sum (map length L) + length r := by - simp [foldr_permutationsAux2, Function.comp_def, length_permutationsAux2, length_flatMap'] + (map length L).sum + length r := by + simp [foldr_permutationsAux2, Function.comp_def, length_permutationsAux2, length_flatMap] -set_option linter.deprecated false in theorem length_foldr_permutationsAux2' (t : α) (ts : List α) (r L : List (List α)) (n) (H : ∀ l ∈ L, length l = n) : length (foldr (fun y r => (permutationsAux2 t ts r y id).2) r L) = n * length L + length r := by - rw [length_foldr_permutationsAux2, (_ : Nat.sum (map length L) = n * length L)] + rw [length_foldr_permutationsAux2, (_ : (map length L).sum = n * length L)] induction' L with l L ih · simp - have sum_map : Nat.sum (map length L) = n * length L := ih fun l m => H l (mem_cons_of_mem _ m) + have sum_map : (map length L).sum = n * length L := ih fun l m => H l (mem_cons_of_mem _ m) have length_l : length l = n := H _ (mem_cons_self _ _) simp [sum_map, length_l, Nat.mul_add, Nat.add_comm, mul_succ] diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean index d88ad52680183..65ef7c84f0abe 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean @@ -99,8 +99,7 @@ theorem GradedAlgebra.lift_ι_eq (i' : ZMod 2) (x' : evenOdd Q i') : · rw [Nat.succ_eq_add_one, add_comm, Nat.cast_add, Nat.cast_one] rfl | zero => - set_option linter.deprecated false in - rw [AlgHom.map_zero] + rw [map_zero] apply Eq.symm apply DFinsupp.single_eq_zero.mpr; rfl | add x y hx hy ihx ihy => diff --git a/Mathlib/LinearAlgebra/Prod.lean b/Mathlib/LinearAlgebra/Prod.lean index 33f885c1bc407..bae77e00fdd50 100644 --- a/Mathlib/LinearAlgebra/Prod.lean +++ b/Mathlib/LinearAlgebra/Prod.lean @@ -830,8 +830,7 @@ variable {N : Type*} [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] open Function -set_option linter.deprecated false - +set_option linter.deprecated false in /-- An auxiliary construction for `tunnel`. The composition of `f`, followed by the isomorphism back to `K`, followed by the inclusion of this submodule back into `M`. -/ @@ -839,11 +838,13 @@ followed by the inclusion of this submodule back into `M`. -/ def tunnelAux (f : M × N →ₗ[R] M) (Kφ : ΣK : Submodule R M, K ≃ₗ[R] M) : M × N →ₗ[R] M := (Kφ.1.subtype.comp Kφ.2.symm.toLinearMap).comp f +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tunnelAux_injective (f : M × N →ₗ[R] M) (i : Injective f) (Kφ : ΣK : Submodule R M, K ≃ₗ[R] M) : Injective (tunnelAux f Kφ) := (Subtype.val_injective.comp Kφ.2.symm.injective).comp i +set_option linter.deprecated false in /-- Auxiliary definition for `tunnel`. -/ @[deprecated "No deprecation message was provided." (since := "2024-06-05")] def tunnel' (f : M × N →ₗ[R] M) (i : Injective f) : ℕ → ΣK : Submodule R M, K ≃ₗ[R] M @@ -853,6 +854,7 @@ def tunnel' (f : M × N →ₗ[R] M) (i : Injective f) : ℕ → ΣK : Submodule ((Submodule.fst R M N).equivMapOfInjective _ (tunnelAux_injective f i (tunnel' f i n))).symm.trans (Submodule.fstEquiv R M N)⟩ +set_option linter.deprecated false in /-- Give an injective map `f : M × N →ₗ[R] M` we can find a nested sequence of submodules all isomorphic to `M`. -/ @@ -865,6 +867,7 @@ def tunnel (f : M × N →ₗ[R] M) (i : Injective f) : ℕ →o (Submodule R M) rw [Submodule.map_comp, Submodule.map_comp] apply Submodule.map_subtype_le⟩ +set_option linter.deprecated false in /-- Give an injective map `f : M × N →ₗ[R] M` we can find a sequence of submodules all isomorphic to `N`. -/ @@ -872,12 +875,14 @@ all isomorphic to `N`. def tailing (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Submodule R M := (Submodule.snd R M N).map (tunnelAux f (tunnel' f i n)) +set_option linter.deprecated false in /-- Each `tailing f i n` is a copy of `N`. -/ @[deprecated "No deprecation message was provided." (since := "2024-06-05")] def tailingLinearEquiv (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailing f i n ≃ₗ[R] N := ((Submodule.snd R M N).equivMapOfInjective _ (tunnelAux_injective f i (tunnel' f i n))).symm.trans (Submodule.sndEquiv R M N) +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailing_le_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailing f i n ≤ OrderDual.ofDual (α := Submodule R M) (tunnel f i n) := by @@ -885,6 +890,7 @@ theorem tailing_le_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : rw [Submodule.map_comp, Submodule.map_comp] apply Submodule.map_subtype_le +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailing_disjoint_tunnel_succ (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Disjoint (tailing f i n) (OrderDual.ofDual (α := Submodule R M) <| tunnel f i (n + 1)) := by @@ -894,6 +900,7 @@ theorem tailing_disjoint_tunnel_succ (f : M × N →ₗ[R] M) (i : Injective f) Submodule.comap_map_eq_of_injective (tunnelAux_injective _ i _), inf_comm, Submodule.fst_inf_snd, Submodule.map_bot] +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailing_sup_tunnel_succ_le_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailing f i n ⊔ (OrderDual.ofDual (α := Submodule R M) <| tunnel f i (n + 1)) ≤ @@ -902,19 +909,23 @@ theorem tailing_sup_tunnel_succ_le_tunnel (f : M × N →ₗ[R] M) (i : Injectiv rw [← Submodule.map_sup, sup_comm, Submodule.fst_sup_snd, Submodule.map_comp, Submodule.map_comp] apply Submodule.map_subtype_le +set_option linter.deprecated false in /-- The supremum of all the copies of `N` found inside the tunnel. -/ @[deprecated "No deprecation message was provided." (since := "2024-06-05")] def tailings (f : M × N →ₗ[R] M) (i : Injective f) : ℕ → Submodule R M := partialSups (tailing f i) +set_option linter.deprecated false in @[simp, deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailings_zero (f : M × N →ₗ[R] M) (i : Injective f) : tailings f i 0 = tailing f i 0 := by simp [tailings] +set_option linter.deprecated false in @[simp, deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailings_succ (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailings f i (n + 1) = tailings f i n ⊔ tailing f i (n + 1) := by simp [tailings] +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailings_disjoint_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Disjoint (tailings f i n) (OrderDual.ofDual (α := Submodule R M) <| tunnel f i (n + 1)) := by @@ -927,6 +938,7 @@ theorem tailings_disjoint_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : · apply Disjoint.mono_right _ ih apply tailing_sup_tunnel_succ_le_tunnel +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailings_disjoint_tailing (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Disjoint (tailings f i n) (tailing f i (n + 1)) := diff --git a/Mathlib/Order/Extension/Well.lean b/Mathlib/Order/Extension/Well.lean index 1b5175e593d0c..3d74ec49658eb 100644 --- a/Mathlib/Order/Extension/Well.lean +++ b/Mathlib/Order/Extension/Well.lean @@ -70,10 +70,9 @@ end IsWellFounded namespace WellFounded -set_option linter.deprecated false - variable (hwf : WellFounded r) +set_option linter.deprecated false in /-- An arbitrary well order on `α` that extends `r`. The construction maps `r` into two well-orders: the first map is `WellFounded.rank`, which is not @@ -89,12 +88,14 @@ noncomputable def wellOrderExtension : LinearOrder α := @LinearOrder.lift' α (Ordinal ×ₗ Cardinal) _ (fun a : α => (hwf.rank a, embeddingToCardinal a)) fun _ _ h => embeddingToCardinal.injective <| congr_arg Prod.snd h +set_option linter.deprecated false in @[deprecated IsWellFounded.wellOrderExtension.isWellFounded_lt (since := "2024-09-07")] instance wellOrderExtension.isWellFounded_lt : IsWellFounded α hwf.wellOrderExtension.lt := ⟨InvImage.wf (fun a : α => (hwf.rank a, embeddingToCardinal a)) <| Ordinal.lt_wf.prod_lex Cardinal.lt_wf⟩ include hwf in +set_option linter.deprecated false in /-- Any well-founded relation can be extended to a well-ordering on that type. -/ @[deprecated IsWellFounded.exists_well_order_ge (since := "2024-09-07")] theorem exists_well_order_ge : ∃ s, r ≤ s ∧ IsWellOrder α s := diff --git a/Mathlib/Order/OmegaCompletePartialOrder.lean b/Mathlib/Order/OmegaCompletePartialOrder.lean index c34eaa969055e..9e76dd904111b 100644 --- a/Mathlib/Order/OmegaCompletePartialOrder.lean +++ b/Mathlib/Order/OmegaCompletePartialOrder.lean @@ -296,8 +296,7 @@ lemma ωScottContinuous.comp (hg : ωScottContinuous g) (hf : ωScottContinuous lemma ωScottContinuous.const {x : β} : ωScottContinuous (Function.const α x) := by simp [ωScottContinuous, ScottContinuousOn, Set.range_nonempty] -set_option linter.deprecated false - +set_option linter.deprecated false in /-- A monotone function `f : α →o β` is continuous if it distributes over ωSup. In order to distinguish it from the (more commonly used) continuity from topology @@ -308,6 +307,7 @@ in Scott topological spaces (not defined here). -/ def Continuous (f : α →o β) : Prop := ∀ c : Chain α, f (ωSup c) = ωSup (c.map f) +set_option linter.deprecated false in /-- `Continuous' f` asserts that `f` is both monotone and continuous. -/ @[deprecated ωScottContinuous (since := "2024-05-29")] def Continuous' (f : α → β) : Prop := @@ -318,6 +318,7 @@ lemma isLUB_of_scottContinuous {c : Chain α} {f : α → β} (hf : ScottContinu IsLUB (Set.range (Chain.map c ⟨f, (ScottContinuous.monotone hf)⟩)) (f (ωSup c)) := ωScottContinuous.isLUB hf.scottContinuousOn +set_option linter.deprecated false in @[deprecated ScottContinuous.ωScottContinuous (since := "2024-05-29")] lemma ScottContinuous.continuous' {f : α → β} (hf : ScottContinuous f) : Continuous' f := by constructor @@ -325,45 +326,55 @@ lemma ScottContinuous.continuous' {f : α → β} (hf : ScottContinuous f) : Con rw [← (ωSup_eq_of_isLUB (isLUB_of_scottContinuous hf))] simp only [OrderHom.coe_mk] +set_option linter.deprecated false in @[deprecated ωScottContinuous.monotone (since := "2024-05-29")] theorem Continuous'.to_monotone {f : α → β} (hf : Continuous' f) : Monotone f := hf.fst +set_option linter.deprecated false in @[deprecated ωScottContinuous.of_monotone_map_ωSup (since := "2024-05-29")] theorem Continuous.of_bundled (f : α → β) (hf : Monotone f) (hf' : Continuous ⟨f, hf⟩) : Continuous' f := ⟨hf, hf'⟩ +set_option linter.deprecated false in @[deprecated ωScottContinuous.of_monotone_map_ωSup (since := "2024-05-29")] theorem Continuous.of_bundled' (f : α →o β) (hf' : Continuous f) : Continuous' f := ⟨f.mono, hf'⟩ +set_option linter.deprecated false in @[deprecated ωScottContinuous_iff_monotone_map_ωSup (since := "2024-05-29")] theorem Continuous'.to_bundled (f : α → β) (hf : Continuous' f) : Continuous ⟨f, hf.to_monotone⟩ := hf.snd +set_option linter.deprecated false in @[simp, norm_cast, deprecated ωScottContinuous_iff_monotone_map_ωSup (since := "2024-05-29")] theorem continuous'_coe : ∀ {f : α →o β}, Continuous' f ↔ Continuous f | ⟨_, hf⟩ => ⟨fun ⟨_, hc⟩ => hc, fun hc => ⟨hf, hc⟩⟩ variable (f : α →o β) (g : β →o γ) +set_option linter.deprecated false in @[deprecated ωScottContinuous.id (since := "2024-05-29")] theorem continuous_id : Continuous (@OrderHom.id α _) := by intro c; rw [c.map_id]; rfl +set_option linter.deprecated false in @[deprecated ωScottContinuous.comp (since := "2024-05-29")] theorem continuous_comp (hfc : Continuous f) (hgc : Continuous g) : Continuous (g.comp f) := by dsimp [Continuous] at *; intro rw [hfc, hgc, Chain.map_comp] +set_option linter.deprecated false in @[deprecated ωScottContinuous.id (since := "2024-05-29")] theorem id_continuous' : Continuous' (@id α) := continuous_id.of_bundled' _ +set_option linter.deprecated false in @[deprecated ωScottContinuous.const (since := "2024-05-29")] theorem continuous_const (x : β) : Continuous (OrderHom.const α x) := fun c => eq_of_forall_ge_iff fun z => by rw [ωSup_le_iff, Chain.map_coe, OrderHom.const_coe_coe]; simp +set_option linter.deprecated false in @[deprecated ωScottContinuous.const (since := "2024-05-29")] theorem const_continuous' (x : β) : Continuous' (Function.const α x) := Continuous.of_bundled' (OrderHom.const α x) (continuous_const x) @@ -484,13 +495,13 @@ lemma ωScottContinuous.of_apply₂ (hf : ∀ a, ωScottContinuous (f · a)) : lemma ωScottContinuous_iff_apply₂ : ωScottContinuous f ↔ ∀ a, ωScottContinuous (f · a) := ⟨ωScottContinuous.apply₂, ωScottContinuous.of_apply₂⟩ -set_option linter.deprecated false - +set_option linter.deprecated false in @[deprecated ωScottContinuous.apply₂ (since := "2024-05-29")] theorem flip₁_continuous' (f : ∀ x : α, γ → β x) (a : α) (hf : Continuous' fun x y => f y x) : Continuous' (f a) := Continuous.of_bundled _ (fun _ _ h => hf.to_monotone h a) fun c => congr_fun (hf.to_bundled _ c) a +set_option linter.deprecated false in @[deprecated ωScottContinuous.of_apply₂ (since := "2024-05-29")] theorem flip₂_continuous' (f : γ → ∀ x, β x) (hf : ∀ x, Continuous' fun g => f g x) : Continuous' f := @@ -571,8 +582,7 @@ lemma ωScottContinuous.top : ωScottContinuous (⊤ : α → β) := lemma ωScottContinuous.bot : ωScottContinuous (⊥ : α → β) := by rw [← sSup_empty]; exact ωScottContinuous.sSup (by simp) -set_option linter.deprecated false - +set_option linter.deprecated false in @[deprecated ωScottContinuous.sSup (since := "2024-05-29")] theorem sSup_continuous (s : Set <| α →o β) (hs : ∀ f ∈ s, Continuous f) : Continuous (sSup s) := by intro c @@ -582,11 +592,13 @@ theorem sSup_continuous (s : Set <| α →o β) (hs : ∀ f ∈ s, Continuous f) simpa (config := { contextual := true }) [ωSup_le_iff, hs _ _ _] using this exact ⟨fun H n f hf => H f hf n, fun H f hf n => H n f hf⟩ +set_option linter.deprecated false in @[deprecated ωScottContinuous.iSup (since := "2024-05-29")] theorem iSup_continuous {ι : Sort*} {f : ι → α →o β} (h : ∀ i, Continuous (f i)) : Continuous (⨆ i, f i) := sSup_continuous _ <| Set.forall_mem_range.2 h +set_option linter.deprecated false in @[deprecated ωScottContinuous.sSup (since := "2024-05-29")] theorem sSup_continuous' (s : Set (α → β)) (hc : ∀ f ∈ s, Continuous' f) : Continuous' (sSup s) := by @@ -596,18 +608,21 @@ theorem sSup_continuous' (s : Set (α → β)) (hc : ∀ f ∈ s, Continuous' f) norm_cast exact iSup_continuous fun f ↦ iSup_continuous fun hf ↦ hc hf +set_option linter.deprecated false in @[deprecated ωScottContinuous.sup (since := "2024-05-29")] theorem sup_continuous {f g : α →o β} (hf : Continuous f) (hg : Continuous g) : Continuous (f ⊔ g) := by rw [← sSup_pair]; apply sSup_continuous rintro f (rfl | rfl | _) <;> assumption +set_option linter.deprecated false in @[deprecated ωScottContinuous.top (since := "2024-05-29")] theorem top_continuous : Continuous (⊤ : α →o β) := by intro c; apply eq_of_forall_ge_iff; intro z simp only [OrderHom.instTopOrderHom_top, OrderHom.const_coe_coe, Function.const, top_le_iff, ωSup_le_iff, Chain.map_coe, Function.comp, forall_const] +set_option linter.deprecated false in @[deprecated ωScottContinuous.bot (since := "2024-05-29")] theorem bot_continuous : Continuous (⊥ : α →o β) := by rw [← sSup_empty] @@ -635,8 +650,7 @@ lemma ωScottContinuous.inf (hf : ωScottContinuous f) (hg : ωScottContinuous g (h (max j i)).imp (le_trans <| hf.monotone <| c.mono <| le_max_left _ _) (le_trans <| hg.monotone <| c.mono <| le_max_right _ _)⟩ -set_option linter.deprecated false - +set_option linter.deprecated false in @[deprecated ωScottContinuous.inf (since := "2024-05-29")] theorem inf_continuous (f g : α →o β) (hf : Continuous f) (hg : Continuous g) : Continuous (f ⊓ g) := by @@ -647,6 +661,7 @@ theorem inf_continuous (f g : α →o β) (hf : Continuous f) (hg : Continuous g (h (max j i)).imp (le_trans <| f.mono <| c.mono <| le_max_left _ _) (le_trans <| g.mono <| c.mono <| le_max_right _ _)⟩ +set_option linter.deprecated false in @[deprecated ωScottContinuous.inf (since := "2024-05-29")] theorem inf_continuous' {f g : α → β} (hf : Continuous' f) (hg : Continuous' g) : Continuous' (f ⊓ g) := @@ -779,8 +794,7 @@ lemma ωScottContinuous.seq {β γ} {f : α → Part (β → γ)} {g : α → Pa simp only [seq_eq_bind_map] exact ωScottContinuous.bind hf <| ωScottContinuous.of_apply₂ fun _ ↦ ωScottContinuous.map hg -set_option linter.deprecated false - +set_option linter.deprecated false in @[deprecated ωScottContinuous.bind (since := "2024-05-29")] theorem bind_continuous' {β γ : Type v} (f : α → Part β) (g : α → β → Part γ) : Continuous' f → Continuous' g → Continuous' fun x => f x >>= g x @@ -788,11 +802,13 @@ theorem bind_continuous' {β γ : Type v} (f : α → Part β) (g : α → β Continuous.of_bundled' (OrderHom.partBind ⟨f, hf⟩ ⟨g, hg⟩) (by intro c; rw [ωSup_bind, ← hf', ← hg']; rfl) +set_option linter.deprecated false in @[deprecated ωScottContinuous.map (since := "2024-05-29")] theorem map_continuous' {β γ : Type v} (f : β → γ) (g : α → Part β) (hg : Continuous' g) : Continuous' fun x => f <$> g x := by simp only [map_eq_bind_pure_comp]; apply bind_continuous' _ _ hg; apply const_continuous' +set_option linter.deprecated false in @[deprecated ωScottContinuous.seq (since := "2024-05-29")] theorem seq_continuous' {β γ : Type v} (f : α → Part (β → γ)) (g : α → Part β) (hf : Continuous' f) (hg : Continuous' g) : Continuous' fun x => f x <*> g x := by @@ -802,8 +818,6 @@ theorem seq_continuous' {β γ : Type v} (f : α → Part (β → γ)) (g : α intro apply map_continuous' _ _ hg -set_option linter.deprecated true - theorem continuous (F : α →𝒄 β) (C : Chain α) : F (ωSup C) = ωSup (C.map F) := F.ωScottContinuous.map_ωSup _ diff --git a/Mathlib/Order/WellFounded.lean b/Mathlib/Order/WellFounded.lean index 0511b5790ae02..17b5c843b5ec7 100644 --- a/Mathlib/Order/WellFounded.lean +++ b/Mathlib/Order/WellFounded.lean @@ -90,9 +90,8 @@ protected theorem lt_sup {r : α → α → Prop} (wf : WellFounded r) {s : Set section deprecated -set_option linter.deprecated false - open Classical in +set_option linter.deprecated false in /-- A successor of an element `x` in a well-founded order is a minimal element `y` such that `x < y` if one exists. Otherwise it is `x` itself. -/ @[deprecated "If you have a linear order, consider defining a `SuccOrder` instance through @@ -100,12 +99,14 @@ open Classical in protected noncomputable def succ {r : α → α → Prop} (wf : WellFounded r) (x : α) : α := if h : ∃ y, r x y then wf.min { y | r x y } h else x +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-25")] protected theorem lt_succ {r : α → α → Prop} (wf : WellFounded r) {x : α} (h : ∃ y, r x y) : r x (wf.succ x) := by rw [WellFounded.succ, dif_pos h] apply min_mem +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-25")] protected theorem lt_succ_iff {r : α → α → Prop} [wo : IsWellOrder α r] {x : α} (h : ∃ y, r x y) (y : α) : r y (wo.wf.succ x) ↔ r y x ∨ y = x := by diff --git a/Mathlib/SetTheory/Cardinal/Aleph.lean b/Mathlib/SetTheory/Cardinal/Aleph.lean index ae9eeee389945..52eabc3d86978 100644 --- a/Mathlib/SetTheory/Cardinal/Aleph.lean +++ b/Mathlib/SetTheory/Cardinal/Aleph.lean @@ -504,7 +504,6 @@ theorem lift_eq_aleph1 {c : Cardinal.{u}} : lift.{v} c = ℵ₁ ↔ c = ℵ₁ : section deprecated -set_option linter.deprecated false set_option linter.docPrime false @[deprecated preAleph (since := "2024-10-22")] @@ -522,6 +521,7 @@ noncomputable alias aleph' := preAleph def alephIdx.initialSeg : @InitialSeg Cardinal Ordinal (· < ·) (· < ·) := @RelEmbedding.collapse Cardinal Ordinal (· < ·) (· < ·) _ Cardinal.ord.orderEmbedding.ltEmbedding +set_option linter.deprecated false in /-- The `aleph'` index function, which gives the ordinal index of a cardinal. (The `aleph'` part is because unlike `aleph` this counts also the finite stages. So `alephIdx n = n`, `alephIdx ℵ₀ = ω`, @@ -533,6 +533,7 @@ def alephIdx.initialSeg : @InitialSeg Cardinal Ordinal (· < ·) (· < ·) := def alephIdx.relIso : @RelIso Cardinal.{u} Ordinal.{u} (· < ·) (· < ·) := aleph'.symm.toRelIsoLT +set_option linter.deprecated false in /-- The `aleph'` index function, which gives the ordinal index of a cardinal. (The `aleph'` part is because unlike `aleph` this counts also the finite stages. So `alephIdx n = n`, `alephIdx ω = ω`, @@ -542,10 +543,12 @@ def alephIdx.relIso : @RelIso Cardinal.{u} Ordinal.{u} (· < ·) (· < ·) := def alephIdx : Cardinal → Ordinal := aleph'.symm +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-08-28")] theorem alephIdx.relIso_coe : (alephIdx.relIso : Cardinal → Ordinal) = alephIdx := rfl +set_option linter.deprecated false in /-- The `aleph'` function gives the cardinals listed by their ordinal index, and is the inverse of `aleph_idx`. `aleph' n = n`, `aleph' ω = ω`, `aleph' (ω + 1) = succ ℵ₀`, etc. @@ -556,55 +559,68 @@ theorem alephIdx.relIso_coe : (alephIdx.relIso : Cardinal → Ordinal) = alephId def Aleph'.relIso := aleph' +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-08-28")] theorem aleph'.relIso_coe : (Aleph'.relIso : Ordinal → Cardinal) = aleph' := rfl +set_option linter.deprecated false in @[deprecated preAleph_lt_preAleph (since := "2024-10-22")] theorem aleph'_lt {o₁ o₂ : Ordinal} : aleph' o₁ < aleph' o₂ ↔ o₁ < o₂ := aleph'.lt_iff_lt +set_option linter.deprecated false in @[deprecated preAleph_le_preAleph (since := "2024-10-22")] theorem aleph'_le {o₁ o₂ : Ordinal} : aleph' o₁ ≤ aleph' o₂ ↔ o₁ ≤ o₂ := aleph'.le_iff_le +set_option linter.deprecated false in @[deprecated preAleph_max (since := "2024-10-22")] theorem aleph'_max (o₁ o₂ : Ordinal) : aleph' (max o₁ o₂) = max (aleph' o₁) (aleph' o₂) := aleph'.monotone.map_max +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-08-28")] theorem aleph'_alephIdx (c : Cardinal) : aleph' c.alephIdx = c := Cardinal.alephIdx.relIso.toEquiv.symm_apply_apply c +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-08-28")] theorem alephIdx_aleph' (o : Ordinal) : (aleph' o).alephIdx = o := Cardinal.alephIdx.relIso.toEquiv.apply_symm_apply o +set_option linter.deprecated false in @[deprecated preAleph_zero (since := "2024-10-22")] theorem aleph'_zero : aleph' 0 = 0 := aleph'.map_bot +set_option linter.deprecated false in @[deprecated preAleph_succ (since := "2024-10-22")] theorem aleph'_succ (o : Ordinal) : aleph' (succ o) = succ (aleph' o) := aleph'.map_succ o +set_option linter.deprecated false in @[deprecated preAleph_nat (since := "2024-10-22")] theorem aleph'_nat : ∀ n : ℕ, aleph' n = n := preAleph_nat +set_option linter.deprecated false in @[deprecated lift_preAleph (since := "2024-10-22")] theorem lift_aleph' (o : Ordinal.{u}) : lift.{v} (aleph' o) = aleph' (Ordinal.lift.{v} o) := lift_preAleph o +set_option linter.deprecated false in @[deprecated preAleph_le_of_isLimit (since := "2024-10-22")] theorem aleph'_le_of_limit {o : Ordinal} (l : o.IsLimit) {c} : aleph' o ≤ c ↔ ∀ o' < o, aleph' o' ≤ c := preAleph_le_of_isLimit l +set_option linter.deprecated false in @[deprecated preAleph_limit (since := "2024-10-22")] theorem aleph'_limit {o : Ordinal} (ho : o.IsLimit) : aleph' o = ⨆ a : Iio o, aleph' a := preAleph_limit ho +set_option linter.deprecated false in @[deprecated preAleph_omega0 (since := "2024-10-22")] theorem aleph'_omega0 : aleph' ω = ℵ₀ := preAleph_omega0 @@ -612,6 +628,7 @@ theorem aleph'_omega0 : aleph' ω = ℵ₀ := @[deprecated "No deprecation message was provided." (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")] def aleph'Equiv : Ordinal ≃ Cardinal := @@ -621,14 +638,17 @@ def aleph'Equiv : Ordinal ≃ Cardinal := theorem aleph_eq_aleph' (o : Ordinal) : ℵ_ o = preAleph (ω + o) := rfl +set_option linter.deprecated false in @[deprecated aleph0_le_preAleph (since := "2024-10-22")] theorem aleph0_le_aleph' {o : Ordinal} : ℵ₀ ≤ aleph' o ↔ ω ≤ o := by rw [← aleph'_omega0, aleph'_le] +set_option linter.deprecated false in @[deprecated preAleph_pos (since := "2024-10-22")] theorem aleph'_pos {o : Ordinal} (ho : 0 < o) : 0 < aleph' o := by rwa [← aleph'_zero, aleph'_lt] +set_option linter.deprecated false in @[deprecated preAleph_isNormal (since := "2024-10-22")] theorem aleph'_isNormal : IsNormal (ord ∘ aleph') := preAleph_isNormal @@ -647,15 +667,18 @@ theorem ord_card_unbounded : Unbounded (· < ·) { b : Ordinal | b.card.ord = b dsimp rw [card_ord], (lt_ord_succ_card a).le⟩⟩ +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-24")] theorem eq_aleph'_of_eq_card_ord {o : Ordinal} (ho : o.card.ord = o) : ∃ a, (aleph' a).ord = o := ⟨aleph'.symm o.card, by simpa using ho⟩ +set_option linter.deprecated false in /-- Infinite ordinals that are cardinals are unbounded. -/ @[deprecated "No deprecation message was provided." (since := "2024-09-24")] theorem ord_card_unbounded' : Unbounded (· < ·) { b : Ordinal | b.card.ord = b ∧ ω ≤ b } := (unbounded_lt_inter_le ω).2 ord_card_unbounded +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-24")] theorem eq_aleph_of_eq_card_ord {o : Ordinal} (ho : o.card.ord = o) (ho' : ω ≤ o) : ∃ a, (ℵ_ a).ord = o := by diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index d4660d1f437be..cb32fdc0b2c53 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -815,19 +815,21 @@ protected theorem isSuccLimit_iff {c : Cardinal} : IsSuccLimit c ↔ c ≠ 0 ∧ section deprecated -set_option linter.deprecated false - +set_option linter.deprecated false in @[deprecated IsSuccLimit.isSuccPrelimit (since := "2024-09-17")] protected theorem IsLimit.isSuccPrelimit {c} (h : IsLimit c) : IsSuccPrelimit c := h.2 +set_option linter.deprecated false in @[deprecated ne_zero_of_isSuccLimit (since := "2024-09-17")] protected theorem IsLimit.ne_zero {c} (h : IsLimit c) : c ≠ 0 := h.1 +set_option linter.deprecated false in @[deprecated IsLimit.isSuccPrelimit (since := "2024-09-05")] alias IsLimit.isSuccLimit := IsLimit.isSuccPrelimit +set_option linter.deprecated false in @[deprecated IsSuccLimit.succ_lt (since := "2024-09-17")] theorem IsLimit.succ_lt {x c} (h : IsLimit c) : x < c → succ x < c := h.isSuccPrelimit.succ_lt @@ -1519,17 +1521,18 @@ theorem aleph0_le_of_isSuccLimit {c : Cardinal} (h : IsSuccLimit c) : ℵ₀ ≤ section deprecated -set_option linter.deprecated false - +set_option linter.deprecated false in @[deprecated isSuccLimit_aleph0 (since := "2024-09-17")] theorem isLimit_aleph0 : IsLimit ℵ₀ := ⟨aleph0_ne_zero, isSuccPrelimit_aleph0⟩ +set_option linter.deprecated false in @[deprecated not_isSuccLimit_natCast (since := "2024-09-17")] lemma not_isLimit_natCast : (n : ℕ) → ¬ IsLimit (n : Cardinal.{u}) | 0, e => e.1 rfl | Nat.succ n, e => Order.not_isSuccPrelimit_succ _ (nat_succ n ▸ e.2) +set_option linter.deprecated false in @[deprecated aleph0_le_of_isSuccLimit (since := "2024-09-17")] theorem IsLimit.aleph0_le {c : Cardinal} (h : IsLimit c) : ℵ₀ ≤ c := by by_contra! h' diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 3563908538d09..cc4daa739e721 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -1464,8 +1464,8 @@ theorem sInf_compl_lt_ord_succ {ι : Type u} (f : ι → Ordinal.{u}) : -- TODO: remove `bsup` in favor of `iSup` in a future refactor. section bsup -set_option linter.deprecated false +set_option linter.deprecated false in private theorem sup_le_sup {ι ι' : Type u} (r : ι → ι → Prop) (r' : ι' → ι' → Prop) [IsWellOrder ι r] [IsWellOrder ι' r'] {o} (ho : type r = o) (ho' : type r' = o) (f : ∀ a < o, Ordinal.{max u v}) : @@ -1480,23 +1480,27 @@ private theorem sup_le_sup {ι ι' : Type u} (r : ι → ι → Prop) (r' : ι' simp_rw [familyOfBFamily', ← hj] apply le_sup +set_option linter.deprecated false in theorem sup_eq_sup {ι ι' : Type u} (r : ι → ι → Prop) (r' : ι' → ι' → Prop) [IsWellOrder ι r] [IsWellOrder ι' r'] {o : Ordinal.{u}} (ho : type r = o) (ho' : type r' = o) (f : ∀ a < o, Ordinal.{max u v}) : sup.{_, v} (familyOfBFamily' r ho f) = sup.{_, v} (familyOfBFamily' r' ho' f) := sup_eq_of_range_eq.{u, u, v} (by simp) +set_option linter.deprecated false in /-- The supremum of a family of ordinals indexed by the set of ordinals less than some `o : Ordinal.{u}`. This is a special case of `sup` over the family provided by `familyOfBFamily`. -/ def bsup (o : Ordinal.{u}) (f : ∀ a < o, Ordinal.{max u v}) : Ordinal.{max u v} := sup.{_, v} (familyOfBFamily o f) +set_option linter.deprecated false in @[simp] theorem sup_eq_bsup {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) : sup.{_, v} (familyOfBFamily o f) = bsup.{_, v} o f := rfl +set_option linter.deprecated false in @[simp] theorem sup_eq_bsup' {o : Ordinal.{u}} {ι} (r : ι → ι → Prop) [IsWellOrder ι r] (ho : type r = o) (f : ∀ a < o, Ordinal.{max u v}) : sup.{_, v} (familyOfBFamily' r ho f) = bsup.{_, v} o f := @@ -1507,6 +1511,7 @@ theorem sSup_eq_bsup {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) : congr rw [range_familyOfBFamily] +set_option linter.deprecated false in @[simp] theorem bsup_eq_sup' {ι : Type u} (r : ι → ι → Prop) [IsWellOrder ι r] (f : ι → Ordinal.{max u v}) : bsup.{_, v} _ (bfamilyOfFamily' r f) = sup.{_, v} f := by @@ -1518,6 +1523,7 @@ theorem bsup_eq_bsup {ι : Type u} (r r' : ι → ι → Prop) [IsWellOrder ι r bsup.{_, v} _ (bfamilyOfFamily' r f) = bsup.{_, v} _ (bfamilyOfFamily' r' f) := by rw [bsup_eq_sup', bsup_eq_sup'] +set_option linter.deprecated false in @[simp] theorem bsup_eq_sup {ι : Type u} (f : ι → Ordinal.{max u v}) : bsup.{_, v} _ (bfamilyOfFamily f) = sup.{_, v} f := @@ -1530,6 +1536,7 @@ theorem bsup_congr {o₁ o₂ : Ordinal.{u}} (f : ∀ a < o₁, Ordinal.{max u v -- Porting note: `rfl` is required. rfl +set_option linter.deprecated false in theorem bsup_le_iff {o f a} : bsup.{u, v} o f ≤ a ↔ ∀ i h, f i h ≤ a := sup_le_iff.trans ⟨fun h i hi => by @@ -1547,6 +1554,7 @@ theorem lt_bsup {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) {a} : a < bsup.{_, v} o f ↔ ∃ i hi, a < f i hi := by simpa only [not_forall, not_le] using not_congr (@bsup_le_iff.{_, v} _ f a) +set_option linter.deprecated false in theorem IsNormal.bsup {f : Ordinal.{max u v} → Ordinal.{max u w}} (H : IsNormal f) {o : Ordinal.{u}} : ∀ (g : ∀ a < o, Ordinal), o ≠ 0 → f (bsup.{_, v} o g) = bsup.{_, w} o fun a h => f (g a h) := @@ -1558,6 +1566,7 @@ theorem lt_bsup_of_ne_bsup {o : Ordinal.{u}} {f : ∀ a < o, Ordinal.{max u v}} (∀ i h, f i h ≠ bsup.{_, v} o f) ↔ ∀ i h, f i h < bsup.{_, v} o f := ⟨fun hf _ _ => lt_of_le_of_ne (le_bsup _ _ _) (hf _ _), fun hf _ _ => ne_of_lt (hf _ _)⟩ +set_option linter.deprecated false in theorem bsup_not_succ_of_ne_bsup {o : Ordinal.{u}} {f : ∀ a < o, Ordinal.{max u v}} (hf : ∀ {i : Ordinal} (h : i < o), f i h ≠ bsup.{_, v} o f) (a) : a < bsup.{_, v} o f → succ a < bsup.{_, v} o f := by @@ -1589,6 +1598,7 @@ theorem bsup_const {o : Ordinal.{u}} (ho : o ≠ 0) (a : Ordinal.{max u v}) : (bsup.{_, v} o fun _ _ => a) = a := le_antisymm (bsup_le fun _ _ => le_rfl) (le_bsup _ 0 (Ordinal.pos_iff_ne_zero.2 ho)) +set_option linter.deprecated false in @[simp] theorem bsup_one (f : ∀ a < (1 : Ordinal), Ordinal) : bsup 1 f = f 0 zero_lt_one := by simp_rw [← sup_eq_bsup, sup_unique, familyOfBFamily, familyOfBFamily', typein_one_toType] @@ -1604,6 +1614,7 @@ theorem bsup_eq_of_brange_eq {o o'} {f : ∀ a < o, Ordinal} {g : ∀ a < o', Or (h : brange o f = brange o' g) : bsup.{u, max v w} o f = bsup.{v, max u w} o' g := (bsup_le_of_brange_subset.{u, v, w} h.le).antisymm (bsup_le_of_brange_subset.{v, u, w} h.ge) +set_option linter.deprecated false in theorem iSup_eq_bsup {o} {f : ∀ a < o, Ordinal} : ⨆ a : Iio o, f a.1 a.2 = bsup o f := by simp_rw [Iio, bsup, sup, iSup, range_familyOfBFamily, brange, range, Subtype.exists, mem_setOf] @@ -1612,17 +1623,19 @@ end bsup -- TODO: bring the lsub API in line with the sSup / iSup API, or deprecate it altogether. section lsub -set_option linter.deprecated false +set_option linter.deprecated false in /-- The least strict upper bound of a family of ordinals. -/ def lsub {ι} (f : ι → Ordinal) : Ordinal := sup (succ ∘ f) +set_option linter.deprecated false in @[simp] theorem sup_eq_lsub {ι : Type u} (f : ι → Ordinal.{max u v}) : sup.{_, v} (succ ∘ f) = lsub.{_, v} f := rfl +set_option linter.deprecated false in theorem lsub_le_iff {ι : Type u} {f : ι → Ordinal.{max u v}} {a} : lsub.{_, v} f ≤ a ↔ ∀ i, f i < a := by convert sup_le_iff.{_, v} (f := succ ∘ f) (a := a) using 2 @@ -1632,6 +1645,7 @@ theorem lsub_le_iff {ι : Type u} {f : ι → Ordinal.{max u v}} {a} : theorem lsub_le {ι} {f : ι → Ordinal} {a} : (∀ i, f i < a) → lsub f ≤ a := lsub_le_iff.2 +set_option linter.deprecated false in theorem lt_lsub {ι} (f : ι → Ordinal) (i) : f i < lsub f := succ_le_iff.1 (le_sup _ i) @@ -1639,19 +1653,23 @@ theorem lt_lsub_iff {ι : Type u} {f : ι → Ordinal.{max u v}} {a} : a < lsub.{_, v} f ↔ ∃ i, a ≤ f i := by simpa only [not_forall, not_lt, not_le] using not_congr (@lsub_le_iff.{_, v} _ f a) +set_option linter.deprecated false in theorem sup_le_lsub {ι : Type u} (f : ι → Ordinal.{max u v}) : sup.{_, v} f ≤ lsub.{_, v} f := sup_le fun i => (lt_lsub f i).le +set_option linter.deprecated false in theorem lsub_le_sup_succ {ι : Type u} (f : ι → Ordinal.{max u v}) : lsub.{_, v} f ≤ succ (sup.{_, v} f) := lsub_le fun i => lt_succ_iff.2 (le_sup f i) +set_option linter.deprecated false in theorem sup_eq_lsub_or_sup_succ_eq_lsub {ι : Type u} (f : ι → Ordinal.{max u v}) : sup.{_, v} f = lsub.{_, v} f ∨ succ (sup.{_, v} f) = lsub.{_, v} f := by cases' eq_or_lt_of_le (sup_le_lsub.{_, v} f) with h h · exact Or.inl h · exact Or.inr ((succ_le_of_lt h).antisymm (lsub_le_sup_succ f)) +set_option linter.deprecated false in theorem sup_succ_le_lsub {ι : Type u} (f : ι → Ordinal.{max u v}) : succ (sup.{_, v} f) ≤ lsub.{_, v} f ↔ ∃ i, f i = sup.{_, v} f := by refine ⟨fun h => ?_, ?_⟩ @@ -1661,10 +1679,12 @@ theorem sup_succ_le_lsub {ι : Type u} (f : ι → Ordinal.{max u v}) : rw [succ_le_iff, ← hf] exact lt_lsub _ _ +set_option linter.deprecated false in theorem sup_succ_eq_lsub {ι : Type u} (f : ι → Ordinal.{max u v}) : succ (sup.{_, v} f) = lsub.{_, v} f ↔ ∃ i, f i = sup.{_, v} f := (lsub_le_sup_succ f).le_iff_eq.symm.trans (sup_succ_le_lsub f) +set_option linter.deprecated false in theorem sup_eq_lsub_iff_succ {ι : Type u} (f : ι → Ordinal.{max u v}) : sup.{_, v} f = lsub.{_, v} f ↔ ∀ a < lsub.{_, v} f, succ a < lsub.{_, v} f := by refine ⟨fun h => ?_, fun hf => le_antisymm (sup_le_lsub f) (lsub_le fun i => ?_)⟩ @@ -1680,6 +1700,7 @@ theorem sup_eq_lsub_iff_succ {ι : Type u} (f : ι → Ordinal.{max u v}) : rw [heq] at this exact this.false +set_option linter.deprecated false in theorem sup_eq_lsub_iff_lt_sup {ι : Type u} (f : ι → Ordinal.{max u v}) : sup.{_, v} f = lsub.{_, v} f ↔ ∀ i, f i < sup.{_, v} f := ⟨fun h i => by @@ -1702,14 +1723,17 @@ theorem lsub_eq_zero_iff {ι : Type u} (f : ι → Ordinal.{max u v}) : rw [h] at this exact this.false +set_option linter.deprecated false in @[simp] theorem lsub_const {ι} [Nonempty ι] (o : Ordinal) : (lsub fun _ : ι => o) = succ o := sup_const (succ o) +set_option linter.deprecated false in @[simp] theorem lsub_unique {ι} [Unique ι] (f : ι → Ordinal) : lsub f = succ (f default) := sup_unique _ +set_option linter.deprecated false in theorem lsub_le_of_range_subset {ι ι'} {f : ι → Ordinal} {g : ι' → Ordinal} (h : Set.range f ⊆ Set.range g) : lsub.{u, max v w} f ≤ lsub.{v, max u w} g := sup_le_of_range_subset.{u, v, w} (by convert Set.image_subset succ h <;> apply Set.range_comp) @@ -1718,6 +1742,7 @@ theorem lsub_eq_of_range_eq {ι ι'} {f : ι → Ordinal} {g : ι' → Ordinal} (h : Set.range f = Set.range g) : lsub.{u, max v w} f = lsub.{v, max u w} g := (lsub_le_of_range_subset.{u, v, w} h.le).antisymm (lsub_le_of_range_subset.{v, u, w} h.ge) +set_option linter.deprecated false in @[simp] theorem lsub_sum {α : Type u} {β : Type v} (f : α ⊕ β → Ordinal) : lsub.{max u v, w} f = @@ -1731,6 +1756,7 @@ theorem lsub_not_mem_range {ι : Type u} (f : ι → Ordinal.{max u v}) : theorem nonempty_compl_range {ι : Type u} (f : ι → Ordinal.{max u v}) : (Set.range f)ᶜ.Nonempty := ⟨_, lsub_not_mem_range.{_, v} f⟩ +set_option linter.deprecated false in @[simp] theorem lsub_typein (o : Ordinal) : lsub.{u, u} (typein (α := o.toType) (· < ·)) = o := (lsub_le.{u, u} typein_lt_self).antisymm @@ -1740,11 +1766,13 @@ theorem lsub_typein (o : Ordinal) : lsub.{u, u} (typein (α := o.toType) (· < conv_rhs at h => rw [← type_lt o] simpa [typein_enum] using lt_lsub.{u, u} (typein (· < ·)) (enum (· < ·) ⟨_, h⟩)) +set_option linter.deprecated false in theorem sup_typein_limit {o : Ordinal} (ho : ∀ a, a < o → succ a < o) : sup.{u, u} (typein ((· < ·) : o.toType → o.toType → Prop)) = o := by -- Porting note: `rwa` → `rw` & `assumption` rw [(sup_eq_lsub_iff_succ.{u, u} (typein (· < ·))).2] <;> rw [lsub_typein o]; assumption +set_option linter.deprecated false in @[simp] theorem sup_typein_succ {o : Ordinal} : sup.{u, u} (typein ((· < ·) : (succ o).toType → (succ o).toType → Prop)) = o := by @@ -1764,7 +1792,6 @@ end lsub -- both of them at once. section blsub -set_option linter.deprecated false /-- The least strict upper bound of a family of ordinals indexed by the set of ordinals less than some `o : Ordinal.{u}`. @@ -1992,6 +2019,7 @@ def blsub₂ (o₁ o₂ : Ordinal) (op : {a : Ordinal} → (a < o₁) → {b : O Ordinal := lsub (fun x : o₁.toType × o₂.toType => op (typein_lt_self x.1) (typein_lt_self x.2)) +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-11")] theorem lt_blsub₂ {o₁ o₂ : Ordinal} (op : {a : Ordinal} → (a < o₁) → {b : Ordinal} → (b < o₂) → Ordinal) {a b : Ordinal} @@ -2003,7 +2031,6 @@ theorem lt_blsub₂ {o₁ o₂ : Ordinal} end blsub section mex -set_option linter.deprecated false /-! ### Minimum excluded ordinals -/ @@ -2013,33 +2040,40 @@ set_option linter.deprecated false def mex {ι : Type u} (f : ι → Ordinal.{max u v}) : Ordinal := sInf (Set.range f)ᶜ +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem mex_not_mem_range {ι : Type u} (f : ι → Ordinal.{max u v}) : mex.{_, v} f ∉ Set.range f := csInf_mem (nonempty_compl_range.{_, v} f) +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem le_mex_of_forall {ι : Type u} {f : ι → Ordinal.{max u v}} {a : Ordinal} (H : ∀ b < a, ∃ i, f i = b) : a ≤ mex.{_, v} f := by by_contra! h exact mex_not_mem_range f (H _ h) +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem ne_mex {ι : Type u} (f : ι → Ordinal.{max u v}) : ∀ i, f i ≠ mex.{_, v} f := by simpa using mex_not_mem_range.{_, v} f +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem mex_le_of_ne {ι} {f : ι → Ordinal} {a} (ha : ∀ i, f i ≠ a) : mex f ≤ a := csInf_le' (by simp [ha]) +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem exists_of_lt_mex {ι} {f : ι → Ordinal} {a} (ha : a < mex f) : ∃ i, f i = a := by by_contra! ha' exact ha.not_le (mex_le_of_ne ha') +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem mex_le_lsub {ι : Type u} (f : ι → Ordinal.{max u v}) : mex.{_, v} f ≤ lsub.{_, v} f := csInf_le' (lsub_not_mem_range f) +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem mex_monotone {α β : Type u} {f : α → Ordinal.{max u v}} {g : β → Ordinal.{max u v}} (h : Set.range f ⊆ Set.range g) : mex.{_, v} f ≤ mex.{_, v} g := by @@ -2048,6 +2082,7 @@ theorem mex_monotone {α β : Type u} {f : α → Ordinal.{max u v}} {g : β → rw [← hj] at hi exact ne_mex g j hi +set_option linter.deprecated false in @[deprecated sInf_compl_lt_ord_succ (since := "2024-09-20")] theorem mex_lt_ord_succ_mk {ι : Type u} (f : ι → Ordinal.{u}) : mex.{_, u} f < (succ #ι).ord := by @@ -2064,6 +2099,7 @@ theorem mex_lt_ord_succ_mk {ι : Type u} (f : ι → Ordinal.{u}) : convert Cardinal.mk_le_of_injective hg rw [Cardinal.mk_ord_toType (succ #ι)] +set_option linter.deprecated false in /-- The minimum excluded ordinal of a family of ordinals indexed by the set of ordinals less than some `o : Ordinal.{u}`. This is a special case of `mex` over the family provided by `familyOfBFamily`. @@ -2073,17 +2109,20 @@ theorem mex_lt_ord_succ_mk {ι : Type u} (f : ι → Ordinal.{u}) : def bmex (o : Ordinal) (f : ∀ a < o, Ordinal) : Ordinal := mex (familyOfBFamily o f) +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem bmex_not_mem_brange {o : Ordinal} (f : ∀ a < o, Ordinal) : bmex o f ∉ brange o f := by rw [← range_familyOfBFamily] apply mex_not_mem_range +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem le_bmex_of_forall {o : Ordinal} (f : ∀ a < o, Ordinal) {a : Ordinal} (H : ∀ b < a, ∃ i hi, f i hi = b) : a ≤ bmex o f := by by_contra! h exact bmex_not_mem_brange f (H _ h) +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem ne_bmex {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) {i} (hi) : f i hi ≠ bmex.{_, v} o f := by @@ -2092,28 +2131,33 @@ theorem ne_bmex {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) {i} (hi) : -- Porting note: `familyOfBFamily_enum` → `typein_enum` rw [typein_enum] +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem bmex_le_of_ne {o : Ordinal} {f : ∀ a < o, Ordinal} {a} (ha : ∀ i hi, f i hi ≠ a) : bmex o f ≤ a := mex_le_of_ne fun _i => ha _ _ +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem exists_of_lt_bmex {o : Ordinal} {f : ∀ a < o, Ordinal} {a} (ha : a < bmex o f) : ∃ i hi, f i hi = a := by cases' exists_of_lt_mex ha with i hi exact ⟨_, typein_lt_self i, hi⟩ +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem bmex_le_blsub {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) : bmex.{_, v} o f ≤ blsub.{_, v} o f := mex_le_lsub _ +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem bmex_monotone {o o' : Ordinal.{u}} {f : ∀ a < o, Ordinal.{max u v}} {g : ∀ a < o', Ordinal.{max u v}} (h : brange o f ⊆ brange o' g) : bmex.{_, v} o f ≤ bmex.{_, v} o' g := mex_monotone (by rwa [range_familyOfBFamily, range_familyOfBFamily]) +set_option linter.deprecated false in @[deprecated sInf_compl_lt_ord_succ (since := "2024-09-20")] theorem bmex_lt_ord_succ_card {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{u}) : bmex.{_, u} o f < (succ o.card).ord := by diff --git a/Mathlib/SetTheory/Ordinal/FixedPoint.lean b/Mathlib/SetTheory/Ordinal/FixedPoint.lean index 84f2f8db36b0d..00a1b45e0c167 100644 --- a/Mathlib/SetTheory/Ordinal/FixedPoint.lean +++ b/Mathlib/SetTheory/Ordinal/FixedPoint.lean @@ -226,8 +226,6 @@ end section -set_option linter.deprecated false - variable {o : Ordinal.{u}} {f : ∀ b < o, Ordinal.{max u v} → Ordinal.{max u v}} /-- The next common fixed point, at least `a`, for a family of normal functions indexed by ordinals. @@ -238,47 +236,56 @@ def nfpBFamily (o : Ordinal.{u}) (f : ∀ b < o, Ordinal.{max u v} → Ordinal.{ Ordinal.{max u v} → Ordinal.{max u v} := nfpFamily (familyOfBFamily o f) +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_eq_nfpFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) : nfpBFamily.{u, v} o f = nfpFamily (familyOfBFamily o f) := rfl +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem foldr_le_nfpBFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) (a l) : List.foldr (familyOfBFamily o f) a l ≤ nfpBFamily.{u, v} o f a := Ordinal.le_iSup _ _ +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem le_nfpBFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) (a) : a ≤ nfpBFamily.{u, v} o f a := Ordinal.le_iSup (fun _ ↦ List.foldr _ a _) [] +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem lt_nfpBFamily {a b} : a < nfpBFamily.{u, v} o f b ↔ ∃ l, a < List.foldr (familyOfBFamily o f) b l := Ordinal.lt_iSup_iff +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_le_iff {o : Ordinal} {f : ∀ b < o, Ordinal → Ordinal} {a b} : nfpBFamily.{u, v} o f a ≤ b ↔ ∀ l, List.foldr (familyOfBFamily o f) a l ≤ b := Ordinal.iSup_le_iff +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_le {o : Ordinal} {f : ∀ b < o, Ordinal → Ordinal} {a b} : (∀ l, List.foldr (familyOfBFamily o f) a l ≤ b) → nfpBFamily.{u, v} o f a ≤ b := Ordinal.iSup_le +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_monotone (hf : ∀ i hi, Monotone (f i hi)) : Monotone (nfpBFamily.{u, v} o f) := nfpFamily_monotone fun _ => hf _ _ +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem apply_lt_nfpBFamily (H : ∀ i hi, IsNormal (f i hi)) {a b} (hb : b < nfpBFamily.{u, v} o f a) (i hi) : f i hi b < nfpBFamily.{u, v} o f a := by rw [← familyOfBFamily_enum o f] apply apply_lt_nfpFamily (fun _ => H _ _) hb +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem apply_lt_nfpBFamily_iff (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a b} : (∀ i hi, f i hi b < nfpBFamily.{u, v} o f a) ↔ b < nfpBFamily.{u, v} o f a := @@ -287,6 +294,7 @@ theorem apply_lt_nfpBFamily_iff (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) refine (apply_lt_nfpFamily_iff ?_).1 fun _ => h _ _ exact fun _ => H _ _, apply_lt_nfpBFamily H⟩ +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_le_apply (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a b} : (∃ i hi, nfpBFamily.{u, v} o f a ≤ f i hi b) ↔ nfpBFamily.{u, v} o f a ≤ b := by @@ -294,11 +302,13 @@ theorem nfpBFamily_le_apply (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a push_neg exact apply_lt_nfpBFamily_iff.{u, v} ho H +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_le_fp (H : ∀ i hi, Monotone (f i hi)) {a b} (ab : a ≤ b) (h : ∀ i hi, f i hi b ≤ b) : nfpBFamily.{u, v} o f a ≤ b := nfpFamily_le_fp (fun _ => H _ _) ab fun _ => h _ _ +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_fp {i hi} (H : IsNormal (f i hi)) (a) : f i hi (nfpBFamily.{u, v} o f a) = nfpBFamily.{u, v} o f a := by @@ -307,6 +317,7 @@ theorem nfpBFamily_fp {i hi} (H : IsNormal (f i hi)) (a) : rw [familyOfBFamily_enum] exact H +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem apply_le_nfpBFamily (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a b} : (∀ i hi, f i hi b ≤ nfpBFamily.{u, v} o f a) ↔ b ≤ nfpBFamily.{u, v} o f a := by @@ -316,10 +327,12 @@ theorem apply_le_nfpBFamily (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a · rw [← nfpBFamily_fp (H i hi)] exact (H i hi).monotone h +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_eq_self {a} (h : ∀ i hi, f i hi a = a) : nfpBFamily.{u, v} o f a = a := nfpFamily_eq_self fun _ => h _ _ +set_option linter.deprecated false in /-- A generalization of the fixed point lemma for normal functions: any family of normal functions has an unbounded set of common fixed points. -/ @[deprecated "No deprecation message was provided." (since := "2024-10-14")] @@ -330,6 +343,7 @@ theorem not_bddAbove_fp_bfamily (H : ∀ i hi, IsNormal (f i hi)) : rw [Set.mem_iInter₂] exact fun i hi ↦ nfpBFamily_fp (H i hi) _ +set_option linter.deprecated false in /-- A generalization of the fixed point lemma for normal functions: any family of normal functions has an unbounded set of common fixed points. -/ @[deprecated not_bddAbove_fp_bfamily (since := "2024-09-20")] @@ -347,11 +361,13 @@ def derivBFamily (o : Ordinal.{u}) (f : ∀ b < o, Ordinal.{max u v} → Ordinal Ordinal.{max u v} → Ordinal.{max u v} := derivFamily (familyOfBFamily o f) +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem derivBFamily_eq_derivFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) : derivBFamily.{u, v} o f = derivFamily (familyOfBFamily o f) := rfl +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem isNormal_derivBFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) : IsNormal (derivBFamily o f) := @@ -360,6 +376,7 @@ theorem isNormal_derivBFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) @[deprecated isNormal_derivBFamily (since := "2024-10-11")] alias derivBFamily_isNormal := isNormal_derivBFamily +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem derivBFamily_fp {i hi} (H : IsNormal (f i hi)) (a : Ordinal) : f i hi (derivBFamily.{u, v} o f a) = derivBFamily.{u, v} o f a := by @@ -368,6 +385,7 @@ theorem derivBFamily_fp {i hi} (H : IsNormal (f i hi)) (a : Ordinal) : rw [familyOfBFamily_enum] exact H +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem le_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : (∀ i hi, f i hi a ≤ a) ↔ ∃ b, derivBFamily.{u, v} o f b = a := by @@ -378,6 +396,7 @@ theorem le_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : apply h · exact fun _ => H _ _ +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem fp_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : (∀ i hi, f i hi a = a) ↔ ∃ b, derivBFamily.{u, v} o f b = a := by @@ -386,6 +405,7 @@ theorem fp_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : rw [← (H i hi).le_iff_eq] exact h i hi +set_option linter.deprecated false in /-- For a family of normal functions, `Ordinal.derivBFamily` enumerates the common fixed points. -/ @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem derivBFamily_eq_enumOrd (H : ∀ i hi, IsNormal (f i hi)) : diff --git a/Mathlib/SetTheory/Ordinal/Rank.lean b/Mathlib/SetTheory/Ordinal/Rank.lean index 1eb486c71b504..fb8b097d185a7 100644 --- a/Mathlib/SetTheory/Ordinal/Rank.lean +++ b/Mathlib/SetTheory/Ordinal/Rank.lean @@ -97,8 +97,6 @@ theorem IsWellFounded.rank_eq_typein (r) [IsWellOrder α r] : rank r = Ordinal.t namespace WellFounded -set_option linter.deprecated false - variable {r : α → α → Prop} (hwf : WellFounded r) /-- The rank of an element `a` under a well-founded relation `r` is defined inductively as the @@ -108,18 +106,22 @@ smallest ordinal greater than the ranks of all elements below it (i.e. elements noncomputable def rank (a : α) : Ordinal.{u} := (hwf.apply a).rank +set_option linter.deprecated false in @[deprecated IsWellFounded.rank_eq (since := "2024-09-07")] theorem rank_eq : hwf.rank a = ⨆ b : { b // r b a }, Order.succ (hwf.rank b) := (hwf.apply a).rank_eq +set_option linter.deprecated false in @[deprecated IsWellFounded.rank_lt_of_rel (since := "2024-09-07")] theorem rank_lt_of_rel (h : r a b) : hwf.rank a < hwf.rank b := Acc.rank_lt_of_rel _ h +set_option linter.deprecated false in @[deprecated WellFoundedLT.rank_strictMono (since := "2024-09-07")] theorem rank_strictMono [Preorder α] [WellFoundedLT α] : StrictMono (rank <| @wellFounded_lt α _ _) := fun _ _ => rank_lt_of_rel _ +set_option linter.deprecated false in @[deprecated WellFoundedGT.rank_strictAnti (since := "2024-09-07")] theorem rank_strictAnti [Preorder α] [WellFoundedGT α] : StrictAnti (rank <| @wellFounded_gt α _ _) := fun _ _ => rank_lt_of_rel wellFounded_gt diff --git a/Mathlib/SetTheory/ZFC/Basic.lean b/Mathlib/SetTheory/ZFC/Basic.lean index a2b3354096bf0..00484448b82b2 100644 --- a/Mathlib/SetTheory/ZFC/Basic.lean +++ b/Mathlib/SetTheory/ZFC/Basic.lean @@ -443,8 +443,6 @@ def embed : PSet.{max (u + 1) v} := theorem lift_mem_embed : ∀ x : PSet.{u}, PSet.Lift.{u, max (u + 1) v} x ∈ embed.{u, v} := fun x => ⟨⟨x⟩, Equiv.rfl⟩ -set_option linter.deprecated false - /-- Function equivalence is defined so that `f ~ g` iff `∀ x y, x ~ y → f x ~ g y`. This extends to equivalence of `n`-ary functions. -/ @[deprecated "No deprecation message was provided." (since := "2024-09-02")] @@ -452,37 +450,44 @@ def Arity.Equiv : ∀ {n}, OfArity PSet.{u} PSet.{u} n → OfArity PSet.{u} PSet | 0, a, b => PSet.Equiv a b | _ + 1, a, b => ∀ x y : PSet, PSet.Equiv x y → Arity.Equiv (a x) (b y) +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-02")] theorem Arity.equiv_const {a : PSet.{u}} : ∀ n, Arity.Equiv (OfArity.const PSet.{u} a n) (OfArity.const PSet.{u} a n) | 0 => Equiv.rfl | _ + 1 => fun _ _ _ => Arity.equiv_const _ +set_option linter.deprecated false in /-- `resp n` is the collection of n-ary functions on `PSet` that respect equivalence, i.e. when the inputs are equivalent the output is as well. -/ @[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Resp (n) := { x : OfArity PSet.{u} PSet.{u} n // Arity.Equiv x x } +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-02")] instance Resp.inhabited {n} : Inhabited (Resp n) := ⟨⟨OfArity.const _ default _, Arity.equiv_const _⟩⟩ +set_option linter.deprecated false in /-- The `n`-ary image of a `(n + 1)`-ary function respecting equivalence as a function respecting equivalence. -/ @[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Resp.f {n} (f : Resp (n + 1)) (x : PSet) : Resp n := ⟨f.1 x, f.2 _ _ <| Equiv.refl x⟩ +set_option linter.deprecated false in /-- Function equivalence for functions respecting equivalence. See `PSet.Arity.Equiv`. -/ @[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Resp.Equiv {n} (a b : Resp n) : Prop := Arity.Equiv a.1 b.1 +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-02"), refl] protected theorem Resp.Equiv.refl {n} (a : Resp n) : Resp.Equiv a a := a.2 +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-02")] protected theorem Resp.Equiv.euc : ∀ {n} {a b c : Resp n}, Resp.Equiv a b → Resp.Equiv c b → Resp.Equiv a c @@ -490,15 +495,18 @@ protected theorem Resp.Equiv.euc : | n + 1, a, b, c, hab, hcb => fun x y h => @Resp.Equiv.euc n (a.f x) (b.f y) (c.f y) (hab _ _ h) (hcb _ _ <| PSet.Equiv.refl y) +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-02"), symm] protected theorem Resp.Equiv.symm {n} {a b : Resp n} : Resp.Equiv a b → Resp.Equiv b a := (Resp.Equiv.refl b).euc +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-02"), trans] protected theorem Resp.Equiv.trans {n} {x y z : Resp n} (h1 : Resp.Equiv x y) (h2 : Resp.Equiv y z) : Resp.Equiv x z := h1.euc h2.symm +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-02")] instance Resp.setoid {n} : Setoid (Resp n) := ⟨Resp.Equiv, Resp.Equiv.refl, Resp.Equiv.symm, Resp.Equiv.trans⟩ @@ -606,10 +614,9 @@ end ZFSet namespace PSet -set_option linter.deprecated false - namespace Resp +set_option linter.deprecated false in /-- Helper function for `PSet.eval`. -/ @[deprecated "No deprecation message was provided." (since := "2024-09-02")] def evalAux : @@ -625,11 +632,13 @@ def evalAux : (@Quotient.ind _ _ fun q => F b q = F c q) fun z => evalAux.2 (Resp.f b z) (Resp.f c z) (h _ _ (PSet.Equiv.refl z))⟩ +set_option linter.deprecated false in /-- An equivalence-respecting function yields an n-ary ZFC set function. -/ @[deprecated "No deprecation message was provided." (since := "2024-09-02")] def eval (n) : Resp n → OfArity ZFSet.{u} ZFSet.{u} n := evalAux.1 +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-02")] theorem eval_val {n f x} : (@eval (n + 1) f : ZFSet → OfArity ZFSet ZFSet n) ⟦x⟧ = eval n (Resp.f f x) := @@ -637,6 +646,7 @@ theorem eval_val {n f x} : end Resp +set_option linter.deprecated false in /-- A set function is "definable" if it is the image of some n-ary pre-set function. This isn't exactly definability, but is useful as a sufficient condition for functions that have a computable image. -/ @@ -647,17 +657,20 @@ class inductive Definable (n) : OfArity ZFSet.{u} ZFSet.{u} n → Type (u + 1) attribute [deprecated "No deprecation message was provided." (since := "2024-09-02"), instance] Definable.mk +set_option linter.deprecated false in /-- The evaluation of a function respecting equivalence is definable, by that same function. -/ @[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Definable.EqMk {n} (f) : ∀ {s : OfArity ZFSet.{u} ZFSet.{u} n} (_ : Resp.eval _ f = s), Definable n s | _, rfl => ⟨f⟩ +set_option linter.deprecated false in /-- Turns a definable function into a function that respects equivalence. -/ @[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Definable.Resp {n} : ∀ (s : OfArity ZFSet.{u} ZFSet.{u} n) [Definable n s], Resp n | _, ⟨f⟩ => f +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-02")] theorem Definable.eq {n} : ∀ (s : OfArity ZFSet.{u} ZFSet.{u} n) [H : Definable n s], (@Definable.Resp n s H).eval _ = s diff --git a/Mathlib/Topology/Order/Basic.lean b/Mathlib/Topology/Order/Basic.lean index 59abac51b6b7c..38a08ce558259 100644 --- a/Mathlib/Topology/Order/Basic.lean +++ b/Mathlib/Topology/Order/Basic.lean @@ -472,9 +472,10 @@ theorem atBot_le_nhds_bot [OrderBot α] : (atBot : Filter α) ≤ 𝓝 ⊥ := by rw [OrderBot.atBot_eq] apply pure_le_nhds +set_option linter.deprecated false in @[deprecated OrderTop.atTop_eq (since := "2024-02-14")] theorem atTop_le_nhds_top [OrderTop α] : (atTop : Filter α) ≤ 𝓝 ⊤ := - set_option linter.deprecated false in @atBot_le_nhds_bot αᵒᵈ _ _ _ + @atBot_le_nhds_bot αᵒᵈ _ _ _ variable (α) From 2f1fb1babaa39ba2b6f891a5599443096bd72cbe Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Tue, 3 Dec 2024 13:00:15 +0000 Subject: [PATCH 491/829] chore: NoZeroSMulDivisors instances for (Add)MonoidAlgebra/Polynomial (#19592) Also generalizes the instance for Finsupp. --- Mathlib/Algebra/MonoidAlgebra/Defs.lean | 8 ++++++++ Mathlib/Algebra/Polynomial/Basic.lean | 5 +++++ Mathlib/Data/Finsupp/Basic.lean | 7 +++---- 3 files changed, 16 insertions(+), 4 deletions(-) diff --git a/Mathlib/Algebra/MonoidAlgebra/Defs.lean b/Mathlib/Algebra/MonoidAlgebra/Defs.lean index 41e3c6d67e134..da1c6d0c0a369 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Defs.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Defs.lean @@ -345,6 +345,10 @@ variable {S : Type*} instance smulZeroClass [Semiring k] [SMulZeroClass R k] : SMulZeroClass R (MonoidAlgebra k G) := Finsupp.smulZeroClass +instance noZeroSMulDivisors [Zero R] [Semiring k] [SMulZeroClass R k] [NoZeroSMulDivisors R k] : + NoZeroSMulDivisors R (MonoidAlgebra k G) := + Finsupp.noZeroSMulDivisors + instance distribSMul [Semiring k] [DistribSMul R k] : DistribSMul R (MonoidAlgebra k G) := Finsupp.distribSMul _ _ @@ -1063,6 +1067,10 @@ section Semiring instance smulZeroClass [Semiring k] [SMulZeroClass R k] : SMulZeroClass R k[G] := Finsupp.smulZeroClass +instance noZeroSMulDivisors [Zero R] [Semiring k] [SMulZeroClass R k] [NoZeroSMulDivisors R k] : + NoZeroSMulDivisors R k[G] := + Finsupp.noZeroSMulDivisors + variable [Semiring k] [AddMonoid G] instance semiring : Semiring k[G] := diff --git a/Mathlib/Algebra/Polynomial/Basic.lean b/Mathlib/Algebra/Polynomial/Basic.lean index db335123a7319..93d9b0870d987 100644 --- a/Mathlib/Algebra/Polynomial/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Basic.lean @@ -135,6 +135,11 @@ instance smulZeroClass {S : Type*} [SMulZeroClass S R] : SMulZeroClass S R[X] wh smul r p := ⟨r • p.toFinsupp⟩ smul_zero a := congr_arg ofFinsupp (smul_zero a) +instance {S : Type*} [Zero S] [SMulZeroClass S R] [NoZeroSMulDivisors S R] : + NoZeroSMulDivisors S R[X] where + eq_zero_or_eq_zero_of_smul_eq_zero eq := + (eq_zero_or_eq_zero_of_smul_eq_zero <| congr_arg toFinsupp eq).imp id (congr_arg ofFinsupp) + -- to avoid a bug in the `ring` tactic instance (priority := 1) pow : Pow R[X] ℕ where pow p n := npowRec n p diff --git a/Mathlib/Data/Finsupp/Basic.lean b/Mathlib/Data/Finsupp/Basic.lean index efd119c83d79e..6ba35e3f0ddf1 100644 --- a/Mathlib/Data/Finsupp/Basic.lean +++ b/Mathlib/Data/Finsupp/Basic.lean @@ -1427,11 +1427,10 @@ theorem sum_smul_index_addMonoidHom [AddMonoid M] [AddCommMonoid N] [DistribSMul {b : R} {h : α → M →+ N} : ((b • g).sum fun a => h a) = g.sum fun i c => h i (b • c) := sum_mapRange_index fun i => (h i).map_zero -instance noZeroSMulDivisors [Semiring R] [AddCommMonoid M] [Module R M] {ι : Type*} +instance noZeroSMulDivisors [Zero R] [Zero M] [SMulZeroClass R M] {ι : Type*} [NoZeroSMulDivisors R M] : NoZeroSMulDivisors R (ι →₀ M) := - ⟨fun h => - or_iff_not_imp_left.mpr fun hc => - Finsupp.ext fun i => (smul_eq_zero.mp (DFunLike.ext_iff.mp h i)).resolve_left hc⟩ + ⟨fun h => or_iff_not_imp_left.mpr fun hc => Finsupp.ext fun i => + (eq_zero_or_eq_zero_of_smul_eq_zero (DFunLike.ext_iff.mp h i)).resolve_left hc⟩ section DistribMulActionSemiHom variable [Monoid R] [AddMonoid M] [AddMonoid N] [DistribMulAction R M] [DistribMulAction R N] From b7fbbb51e646ab635cc81c18a1ec2c4fda633811 Mon Sep 17 00:00:00 2001 From: Scott Carnahan <128885296+ScottCarnahan@users.noreply.github.com> Date: Tue, 3 Dec 2024 14:07:42 +0000 Subject: [PATCH 492/829] feat(LinearAlgebra/RootSystem/Finite): Coxeter weights are bounded above by 4 (#19343) This PR adds a version of the Cauchy-Schwarz inequality for positive semidefinite bilinear forms on modules over linearly ordered commutative rings, and applies it to restrict Coxeter weights for root systems. Co-authored-by: Oliver Nash --- Mathlib/LinearAlgebra/RootSystem/Basic.lean | 9 +- Mathlib/LinearAlgebra/RootSystem/Defs.lean | 6 ++ .../RootSystem/Finite/CanonicalBilinear.lean | 61 +++++++++++- .../RootSystem/Finite/Nondegenerate.lean | 6 -- Mathlib/LinearAlgebra/SesquilinearForm.lean | 93 +++++++++++++++++++ 5 files changed, 166 insertions(+), 9 deletions(-) diff --git a/Mathlib/LinearAlgebra/RootSystem/Basic.lean b/Mathlib/LinearAlgebra/RootSystem/Basic.lean index 5a32efb2dc700..ba9290e8307a8 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Basic.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Basic.lean @@ -73,7 +73,7 @@ protected def equiv_of_mapsTo : end reflection_perm -lemma infinite_of_linearly_independent_coxeterWeight_four [CharZero R] [NoZeroSMulDivisors ℤ M] +lemma infinite_of_linearIndependent_coxeterWeight_four [CharZero R] [NoZeroSMulDivisors ℤ M] (P : RootPairing ι R M N) (i j : ι) (hl : LinearIndependent R ![P.root i, P.root j]) (hc : P.coxeterWeight i j = 4) : Infinite ι := by refine (infinite_range_iff (Embedding.injective P.root)).mp (Infinite.mono ?_ @@ -94,6 +94,13 @@ lemma infinite_of_linearly_independent_coxeterWeight_four [CharZero R] [NoZeroSM variable [Finite ι] (P : RootPairing ι R M N) (i j : ι) +lemma coxeterWeight_ne_four_of_linearIndependent [CharZero R] [NoZeroSMulDivisors ℤ M] + (hl : LinearIndependent R ![P.root i, P.root j]) : + P.coxeterWeight i j ≠ 4 := by + intro contra + have := P.infinite_of_linearIndependent_coxeterWeight_four i j hl contra + exact not_finite ι + /-- Even though the roots may not span, coroots are distinguished by their pairing with the roots. The proof depends crucially on the fact that there are finitely-many roots. diff --git a/Mathlib/LinearAlgebra/RootSystem/Defs.lean b/Mathlib/LinearAlgebra/RootSystem/Defs.lean index 3970a7e11cf75..66e76b9408a4a 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Defs.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Defs.lean @@ -541,6 +541,12 @@ def coxeterWeight : R := pairing P i j * pairing P j i lemma coxeterWeight_swap : coxeterWeight P i j = coxeterWeight P j i := by simp only [coxeterWeight, mul_comm] +lemma exists_int_eq_coxeterWeight (h : P.IsCrystallographic) (i j : ι) : + ∃ z : ℤ, P.coxeterWeight i j = z := by + obtain ⟨a, ha⟩ := P.isCrystallographic_iff.mp h i j + obtain ⟨b, hb⟩ := P.isCrystallographic_iff.mp h j i + exact ⟨a * b, by simp [coxeterWeight, ha, hb]⟩ + /-- Two roots are orthogonal when they are fixed by each others' reflections. -/ def IsOrthogonal : Prop := pairing P i j = 0 ∧ pairing P j i = 0 diff --git a/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean b/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean index 9917885d01294..3b1ebfefe84f2 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean @@ -3,9 +3,8 @@ 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.Order.BigOperators.Ring.Finset import Mathlib.Algebra.Ring.SumsOfSquares -import Mathlib.LinearAlgebra.RootSystem.Defs +import Mathlib.LinearAlgebra.RootSystem.RootPositive /-! # The canonical bilinear form on a finite root pairing @@ -98,6 +97,15 @@ lemma corootForm_apply_apply (x y : N) : P.CorootForm x y = ∑ i, P.root' i x * P.root' i y := by simp [CorootForm] +lemma toPerfectPairing_apply_apply_Polarization (x y : M) : + P.toPerfectPairing y (P.Polarization x) = P.RootForm x y := by + simp [RootForm] + +lemma toPerfectPairing_apply_CoPolarization (x : N) : + P.toPerfectPairing (P.CoPolarization x) = P.CorootForm x := by + ext y + exact P.flip.toPerfectPairing_apply_apply_Polarization x y + lemma flip_comp_polarization_eq_rootForm : P.flip.toLin ∘ₗ P.Polarization = P.RootForm := by ext; simp [rootForm_apply_apply, RootPairing.flip] @@ -150,6 +158,23 @@ lemma rootForm_self_smul_coroot (i : ι) : rw [Finset.sum_smul, add_neg_eq_zero.mpr rfl] exact sub_eq_zero_of_eq rfl +lemma four_smul_rootForm_sq_eq_coxeterWeight_smul (i j : ι) : + 4 • (P.RootForm (P.root i) (P.root j)) ^ 2 = P.coxeterWeight i j • + (P.RootForm (P.root i) (P.root i) * P.RootForm (P.root j) (P.root j)) := by + have hij : 4 • (P.RootForm (P.root i)) (P.root j) = + 2 • P.toPerfectPairing (P.root j) (2 • P.Polarization (P.root i)) := by + rw [← toPerfectPairing_apply_apply_Polarization, LinearMap.map_smul_of_tower, ← smul_assoc, + Nat.nsmul_eq_mul] + have hji : 2 • (P.RootForm (P.root i)) (P.root j) = + P.toPerfectPairing (P.root i) (2 • P.Polarization (P.root j)) := by + rw [show (P.RootForm (P.root i)) (P.root j) = (P.RootForm (P.root j)) (P.root i) by + apply rootForm_symmetric, ← toPerfectPairing_apply_apply_Polarization, + LinearMap.map_smul_of_tower] + rw [sq, nsmul_eq_mul, ← mul_assoc, ← nsmul_eq_mul, hij, ← rootForm_self_smul_coroot, + smul_mul_assoc 2, ← mul_smul_comm, hji, ← rootForm_self_smul_coroot, map_smul, ← pairing, + map_smul, ← pairing, smul_eq_mul, smul_eq_mul, smul_eq_mul, coxeterWeight] + ring + 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 @@ -229,6 +254,38 @@ lemma rootForm_root_self_pos (j : ι) : use j simp +/-- SGA3 XXI Prop. 2.3.1 -/ +lemma coxeterWeight_le_four (i j : ι) : P.coxeterWeight i j ≤ 4 := by + set li := P.RootForm (P.root i) (P.root i) + set lj := P.RootForm (P.root j) (P.root j) + set lij := P.RootForm (P.root i) (P.root j) + have hi := P.rootForm_root_self_pos i + have hj := P.rootForm_root_self_pos j + have cs : 4 * lij ^ 2 ≤ 4 * (li * lj) := by + rw [mul_le_mul_left four_pos] + exact LinearMap.BilinForm.apply_sq_le_of_symm P.RootForm P.rootForm_self_non_neg + P.rootForm_symmetric (P.root i) (P.root j) + have key : 4 • lij ^ 2 = _ • (li * lj) := P.four_smul_rootForm_sq_eq_coxeterWeight_smul i j + simp only [nsmul_eq_mul, smul_eq_mul, Nat.cast_ofNat] at key + rwa [key, mul_le_mul_right (by positivity)] at cs + +instance instIsRootPositiveRootForm : IsRootPositive P P.RootForm where + zero_lt_apply_root i := P.rootForm_root_self_pos i + symm := P.rootForm_symmetric + apply_reflection_eq := P.rootForm_reflection_reflection_apply + +lemma coxeterWeight_mem_set_of_isCrystallographic (i j : ι) (hP : P.IsCrystallographic) : + P.coxeterWeight i j ∈ ({0, 1, 2, 3, 4} : Set R) := by + obtain ⟨n, hcn⟩ : ∃ n : ℕ, P.coxeterWeight i j = n := by + obtain ⟨z, hz⟩ := P.exists_int_eq_coxeterWeight hP i j + have hz₀ : 0 ≤ z := by simpa [hz] using P.coxeterWeight_non_neg P.RootForm i j + obtain ⟨n, rfl⟩ := Int.eq_ofNat_of_zero_le hz₀ + exact ⟨n, by simp [hz]⟩ + have : P.coxeterWeight i j ≤ 4 := P.coxeterWeight_le_four i j + simp only [hcn, mem_insert_iff, mem_singleton_iff] at this ⊢ + norm_cast at this ⊢ + omega + lemma prod_rootForm_root_self_pos : 0 < ∏ i, P.RootForm (P.root i) (P.root i) := Finset.prod_pos fun i _ => rootForm_root_self_pos P i diff --git a/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean b/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean index 821094cd58377..65fb8e8be87a7 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean @@ -7,7 +7,6 @@ import Mathlib.LinearAlgebra.BilinearForm.Basic import Mathlib.LinearAlgebra.Dimension.Localization import Mathlib.LinearAlgebra.QuadraticForm.Basic import Mathlib.LinearAlgebra.RootSystem.Finite.CanonicalBilinear -import Mathlib.LinearAlgebra.RootSystem.RootPositive /-! # Nondegeneracy of the polarization on a finite root pairing @@ -55,11 +54,6 @@ variable {ι R M N : Type*} variable [Fintype ι] [LinearOrderedCommRing R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] (P : RootPairing ι R M N) -lemma rootForm_rootPositive : IsRootPositive P P.RootForm where - zero_lt_apply_root i := P.rootForm_root_self_pos i - symm := P.rootForm_symmetric - apply_reflection_eq := P.rootForm_reflection_reflection_apply - @[simp] lemma finrank_rootSpan_map_polarization_eq_finrank_corootSpan : finrank R (P.rootSpan.map P.Polarization) = finrank R P.corootSpan := by diff --git a/Mathlib/LinearAlgebra/SesquilinearForm.lean b/Mathlib/LinearAlgebra/SesquilinearForm.lean index 99617d933657d..f4f55354935c4 100644 --- a/Mathlib/LinearAlgebra/SesquilinearForm.lean +++ b/Mathlib/LinearAlgebra/SesquilinearForm.lean @@ -798,4 +798,97 @@ end CommRing end Nondegenerate +namespace BilinForm + +lemma apply_smul_sub_smul_sub_eq [CommRing R] [AddCommGroup M] [Module R M] + (B : LinearMap.BilinForm R M) (x y : M) : + B ((B x y) • x - (B x x) • y) ((B x y) • x - (B x x) • y) = + (B x x) * ((B x x) * (B y y) - (B x y) * (B y x)) := by + simp only [map_sub, map_smul, sub_apply, smul_apply, smul_eq_mul, mul_sub, + mul_comm (B x y) (B x x), mul_left_comm (B x y) (B x x)] + abel + +variable [LinearOrderedCommRing R] [AddCommGroup M] [Module R M] (B : LinearMap.BilinForm R M) + +/-- The **Cauchy-Schwarz inequality** for positive semidefinite forms. -/ +lemma apply_mul_apply_le_of_forall_zero_le (hs : ∀ x, 0 ≤ B x x) (x y : M) : + (B x y) * (B y x) ≤ (B x x) * (B y y) := by + have aux (x y : M) : 0 ≤ (B x x) * ((B x x) * (B y y) - (B x y) * (B y x)) := by + rw [← apply_smul_sub_smul_sub_eq B x y] + exact hs (B x y • x - B x x • y) + rcases lt_or_le 0 (B x x) with hx | hx + · exact sub_nonneg.mp <| nonneg_of_mul_nonneg_right (aux x y) hx + · replace hx : B x x = 0 := le_antisymm hx (hs x) + rcases lt_or_le 0 (B y y) with hy | hy + · rw [mul_comm (B x y), mul_comm (B x x)] + exact sub_nonneg.mp <| nonneg_of_mul_nonneg_right (aux y x) hy + · replace hy : B y y = 0 := le_antisymm hy (hs y) + suffices B x y = - B y x by simpa [this, hx, hy] using mul_self_nonneg (B y x) + rw [eq_neg_iff_add_eq_zero] + apply le_antisymm + · simpa [hx, hy, le_neg_iff_add_nonpos_left] using hs (x - y) + · simpa [hx, hy] using hs (x + y) + +/-- The **Cauchy-Schwarz inequality** for positive semidefinite symmetric forms. -/ +lemma apply_sq_le_of_symm (hs : ∀ x, 0 ≤ B x x) (hB : B.IsSymm) (x y : M) : + (B x y) ^ 2 ≤ (B x x) * (B y y) := by + rw [show (B x y) ^ 2 = (B x y) * (B y x) by rw [sq, ← hB, RingHom.id_apply]] + exact apply_mul_apply_le_of_forall_zero_le B hs x y + +/-- The equality case of **Cauchy-Schwarz**. -/ +lemma not_linearIndependent_of_apply_mul_apply_eq (hp : ∀ x, x ≠ 0 → 0 < B x x) + (x y : M) (he : (B x y) * (B y x) = (B x x) * (B y y)) : + ¬ LinearIndependent R ![x, y] := by + have hz : (B x y) • x - (B x x) • y = 0 := by + by_contra hc + exact (ne_of_lt (hp ((B x) y • x - (B x) x • y) hc)).symm <| + (apply_smul_sub_smul_sub_eq B x y).symm ▸ (mul_eq_zero_of_right ((B x) x) + (sub_eq_zero_of_eq he.symm)) + by_contra hL + by_cases hx : x = 0 + · simpa [hx] using LinearIndependent.ne_zero 0 hL + · have h := sub_eq_zero.mpr (sub_eq_zero.mp hz).symm + rw [sub_eq_add_neg, ← neg_smul, add_comm] at h + exact (Ne.symm (ne_of_lt (hp x hx))) (LinearIndependent.eq_zero_of_pair hL h).2 + +/-- Strict **Cauchy-Schwarz** is equivalent to linear independence for positive definite forms. -/ +lemma apply_mul_apply_lt_iff_linearIndependent [NoZeroSMulDivisors R M] + (hp : ∀ x, x ≠ 0 → 0 < B x x) (x y : M) : + (B x y) * (B y x) < (B x x) * (B y y) ↔ LinearIndependent R ![x, y] := by + have hle : ∀ z, 0 ≤ B z z := by + intro z + by_cases hz : z = 0; simp [hz] + exact le_of_lt (hp z hz) + constructor + · contrapose! + intro h + rw [LinearIndependent.pair_iff] at h + push_neg at h + obtain ⟨r, s, hl, h0⟩ := h + by_cases hr : r = 0; · simp_all + by_cases hs : s = 0; · simp_all + suffices + (B (r • x) (r • x)) * (B (s • y) (s • y)) = (B (r • x) (s • y)) * (B (s • y) (r • x)) by + simp only [map_smul, smul_apply, smul_eq_mul] at this + rw [show r * (r * (B x) x) * (s * (s * (B y) y)) = (r * r * s * s) * ((B x) x * (B y) y) by + ring, show s * (r * (B x) y) * (r * (s * (B y) x)) = (r * r * s * s) * ((B x) y * (B y) x) + by ring] at this + have hrs : r * r * s * s ≠ 0 := by simp [hr, hs] + exact le_of_eq <| mul_right_injective₀ hrs this + simp [show s • y = - r • x by rwa [neg_smul, ← add_eq_zero_iff_eq_neg']] + · contrapose! + intro h + refine not_linearIndependent_of_apply_mul_apply_eq B hp x y (le_antisymm + (apply_mul_apply_le_of_forall_zero_le B hle x y) h) + +/-- Strict **Cauchy-Schwarz** is equivalent to linear independence for positive definite symmetric +forms. -/ +lemma apply_sq_lt_iff_linearIndependent_of_symm [NoZeroSMulDivisors R M] + (hp : ∀ x, x ≠ 0 → 0 < B x x) (hB: B.IsSymm) (x y : M) : + (B x y) ^ 2 < (B x x) * (B y y) ↔ LinearIndependent R ![x, y] := by + rw [show (B x y) ^ 2 = (B x y) * (B y x) by rw [sq, ← hB, RingHom.id_apply]] + exact apply_mul_apply_lt_iff_linearIndependent B hp x y + +end BilinForm + end LinearMap From b10b73014f9af59ab87454a466bd9b1a631f7ffe Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Tue, 3 Dec 2024 14:22:20 +0000 Subject: [PATCH 493/829] feat(Algebra/Module): presentation of the restriction of scalars of a module (#18389) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Given a morphism of rings `A → B` and a `B`-module `M`, we obtain a presentation of `M` as a `A`-module from a presentation of `M` as `B`-module, a presentation of `B` as a `A`-module (and some additional data). Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- Mathlib.lean | 1 + .../Module/Presentation/RestrictScalars.lean | 66 +++++++++++++++++++ 2 files changed, 67 insertions(+) create mode 100644 Mathlib/Algebra/Module/Presentation/RestrictScalars.lean diff --git a/Mathlib.lean b/Mathlib.lean index 42c0594b705cc..d9fc3061091cf 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -554,6 +554,7 @@ import Mathlib.Algebra.Module.Presentation.Differentials import Mathlib.Algebra.Module.Presentation.DirectSum import Mathlib.Algebra.Module.Presentation.Finite import Mathlib.Algebra.Module.Presentation.Free +import Mathlib.Algebra.Module.Presentation.RestrictScalars import Mathlib.Algebra.Module.Presentation.Tautological import Mathlib.Algebra.Module.Prod import Mathlib.Algebra.Module.Projective diff --git a/Mathlib/Algebra/Module/Presentation/RestrictScalars.lean b/Mathlib/Algebra/Module/Presentation/RestrictScalars.lean new file mode 100644 index 0000000000000..108320b3d56d8 --- /dev/null +++ b/Mathlib/Algebra/Module/Presentation/RestrictScalars.lean @@ -0,0 +1,66 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.Algebra.Module.Presentation.DirectSum +import Mathlib.Algebra.Module.Presentation.Cokernel + +/-! +# Presentation of the restriction of scalars of a module + +Given a morphism of rings `A → B` and a `B`-module `M`, we obtain a presentation +of `M` as a `A`-module from a presentation of `M` as `B`-module, +a presentation of `B` as a `A`-module (and some additional data). + +## TODO +* deduce that if `B` is a finitely presented as an `A`-module and `M` is +finitely presented as an `B`-module, then `M` is finitely presented as an `A`-module + +-/ + +namespace Module + +variable {B : Type*} [Ring B] {M : Type*} [AddCommGroup M] [Module B M] + [DecidableEq B] + (presM : Presentation B M) [DecidableEq presM.G] + {A : Type*} [CommRing A] [Algebra A B] [Module A M] [IsScalarTower A B M] + (presB : Presentation A B) + +namespace Presentation + +/-- The additional data that is necessary in order to obtain a presentation +of the restriction of scalars of a module. -/ +abbrev RestrictScalarsData : Type _ := + (presB.finsupp presM.G).CokernelData + (LinearMap.restrictScalars A presM.map) + (fun (⟨g, g'⟩ : presB.G × presM.R) ↦ presB.var g • Finsupp.single g' (1 : B)) + +variable (data : presM.RestrictScalarsData presB) + +/-- A presentation of the restriction of scalars from `B` to `A` of a `B`-module `M`, +given a presentation of `M` as a `B`-module, a presentation of `B` as an `A`-module, +and an additional data. -/ +noncomputable def restrictScalars : Presentation A M := + ofExact (g := LinearMap.restrictScalars A presM.π) (presB.finsupp presM.G) data + presM.exact presM.surjective_π (by + ext v + dsimp + simp only [Submodule.mem_top, iff_true] + apply Finsupp.induction + · simp + · intro r b w _ _ hw + refine Submodule.add_mem _ ?_ hw + obtain ⟨β, rfl⟩ := presB.surjective_π b + apply Finsupp.induction (p := fun β ↦ Finsupp.single r (presB.π β) ∈ _) + · simp + · intro g a f _ _ hf + rw [map_add, Finsupp.single_add] + refine Submodule.add_mem _ ?_ hf + rw [← Finsupp.smul_single_one, ← Finsupp.smul_single_one, + map_smul, Relations.Solution.π_single, smul_assoc] + exact Submodule.smul_mem _ _ (Submodule.subset_span ⟨⟨g, r⟩, rfl⟩)) + +end Presentation + +end Module From ac7c96fc6d0642784c03699b2afcbf3cb183140e Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Tue, 3 Dec 2024 14:22:21 +0000 Subject: [PATCH 494/829] chore(Algebra/Category/ModuleCat): rename `ModuleCat.asHom` to `ModuleCat.ofHom` (#19705) As discussed in the comments for #19511. It seems the choice between `ofHom` and `asHom` is somewhat inconsistent in concrete categories, but `ofHom` wins. In particular, we want to be consistent with the newly refactored `AlgebraCat.ofHom` and the constructor for objects `ModuleCat.of`. The choice between `asHom` and `ofHom` was only made a few months ago: #17476. It's a bit unfortunate to go back on something that was deprecated so recently, but I think the result is worth the pain. I have not touched `asHomLeft` and `asHomRight` since #19511 will replace these with `ofHom` everywhere anyway. --- .../Algebra/Category/AlgebraCat/Basic.lean | 4 +-- .../CoalgebraCat/ComonEquivalence.lean | 6 ++-- .../Algebra/Category/FGModuleCat/Basic.lean | 4 +-- Mathlib/Algebra/Category/ModuleCat/Basic.lean | 14 ++++---- .../Category/ModuleCat/Biproducts.lean | 8 ++--- .../Algebra/Category/ModuleCat/Images.lean | 4 +-- .../Algebra/Category/ModuleCat/Injective.lean | 4 +-- .../Algebra/Category/ModuleCat/Kernels.lean | 8 ++--- .../Homology/ShortComplex/ModuleCat.lean | 4 +-- Mathlib/CategoryTheory/Linear/Yoneda.lean | 10 +++--- .../Preadditive/Yoneda/Basic.lean | 4 +-- .../GroupCohomology/LowDegree.lean | 16 ++++----- Mathlib/RepresentationTheory/Rep.lean | 2 +- .../AdicCompletion/AsTensorProduct.lean | 34 +++++++++---------- .../RingTheory/Coalgebra/TensorProduct.lean | 2 +- Mathlib/RingTheory/Flat/CategoryTheory.lean | 4 +-- .../Topology/Category/Profinite/Nobeling.lean | 12 +++---- 17 files changed, 70 insertions(+), 70 deletions(-) diff --git a/Mathlib/Algebra/Category/AlgebraCat/Basic.lean b/Mathlib/Algebra/Category/AlgebraCat/Basic.lean index 1053dac09d757..1f4d18462f038 100644 --- a/Mathlib/Algebra/Category/AlgebraCat/Basic.lean +++ b/Mathlib/Algebra/Category/AlgebraCat/Basic.lean @@ -153,7 +153,7 @@ instance hasForgetToRing : HasForget₂ (AlgebraCat.{v} R) RingCat.{v} where instance hasForgetToModule : HasForget₂ (AlgebraCat.{v} R) (ModuleCat.{v} R) where forget₂ := { obj := fun M => ModuleCat.of R M - map := fun f => ModuleCat.asHom f.hom.toLinearMap } + map := fun f => ModuleCat.ofHom f.hom.toLinearMap } @[simp] lemma forget₂_module_obj (X : AlgebraCat.{v} R) : @@ -162,7 +162,7 @@ lemma forget₂_module_obj (X : AlgebraCat.{v} R) : @[simp] lemma forget₂_module_map {X Y : AlgebraCat.{v} R} (f : X ⟶ Y) : - (forget₂ (AlgebraCat.{v} R) (ModuleCat.{v} R)).map f = ModuleCat.asHom f.hom.toLinearMap := + (forget₂ (AlgebraCat.{v} R) (ModuleCat.{v} R)).map f = ModuleCat.ofHom f.hom.toLinearMap := rfl variable {R} in diff --git a/Mathlib/Algebra/Category/CoalgebraCat/ComonEquivalence.lean b/Mathlib/Algebra/Category/CoalgebraCat/ComonEquivalence.lean index c150eabdf07e7..322fce71e773d 100644 --- a/Mathlib/Algebra/Category/CoalgebraCat/ComonEquivalence.lean +++ b/Mathlib/Algebra/Category/CoalgebraCat/ComonEquivalence.lean @@ -38,8 +38,8 @@ variable {R : Type u} [CommRing R] /-- An `R`-coalgebra is a comonoid object in the category of `R`-modules. -/ @[simps] def toComonObj (X : CoalgebraCat R) : Comon_ (ModuleCat R) where X := ModuleCat.of R X - counit := ModuleCat.asHom Coalgebra.counit - comul := ModuleCat.asHom Coalgebra.comul + counit := ModuleCat.ofHom Coalgebra.counit + comul := ModuleCat.ofHom Coalgebra.comul counit_comul := by simpa only [ModuleCat.of_coe] using Coalgebra.rTensor_counit_comp_comul comul_counit := by simpa only [ModuleCat.of_coe] using Coalgebra.lTensor_counit_comp_comul comul_assoc := by simp_rw [ModuleCat.of_coe]; exact Coalgebra.coassoc.symm @@ -50,7 +50,7 @@ variable (R) in def toComon : CoalgebraCat R ⥤ Comon_ (ModuleCat R) where obj X := toComonObj X map f := - { hom := ModuleCat.asHom f.1 + { hom := ModuleCat.ofHom f.1 hom_counit := f.1.counit_comp hom_comul := f.1.map_comp_comul.symm } diff --git a/Mathlib/Algebra/Category/FGModuleCat/Basic.lean b/Mathlib/Algebra/Category/FGModuleCat/Basic.lean index e6577b9339afb..613eb1a083c3f 100644 --- a/Mathlib/Algebra/Category/FGModuleCat/Basic.lean +++ b/Mathlib/Algebra/Category/FGModuleCat/Basic.lean @@ -267,8 +267,8 @@ end FGModuleCat @[simp] theorem LinearMap.comp_id_fgModuleCat {R} [Ring R] {G : FGModuleCat.{u} R} {H : Type u} [AddCommGroup H] [Module R H] (f : G →ₗ[R] H) : f.comp (𝟙 G) = f := - Category.id_comp (ModuleCat.asHom f) + Category.id_comp (ModuleCat.ofHom f) @[simp] theorem LinearMap.id_fgModuleCat_comp {R} [Ring R] {G : Type u} [AddCommGroup G] [Module R G] {H : FGModuleCat.{u} R} (f : G →ₗ[R] H) : LinearMap.comp (𝟙 H) f = f := - Category.comp_id (ModuleCat.asHom f) + Category.comp_id (ModuleCat.ofHom f) diff --git a/Mathlib/Algebra/Category/ModuleCat/Basic.lean b/Mathlib/Algebra/Category/ModuleCat/Basic.lean index 417c5c7e12758..5a22d8f540260 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Basic.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Basic.lean @@ -209,21 +209,21 @@ variable {X₁ X₂ : Type v} open ModuleCat /-- Reinterpreting a linear map in the category of `R`-modules. -/ -def ModuleCat.asHom [AddCommGroup X₁] [Module R X₁] [AddCommGroup X₂] [Module R X₂] : +def ModuleCat.ofHom [AddCommGroup X₁] [Module R X₁] [AddCommGroup X₂] [Module R X₂] : (X₁ →ₗ[R] X₂) → (ModuleCat.of R X₁ ⟶ ModuleCat.of R X₂) := id -@[deprecated (since := "2024-10-06")] alias ModuleCat.ofHom := ModuleCat.asHom +@[deprecated (since := "2024-12-03")] alias ModuleCat.asHom := ModuleCat.ofHom /-- Reinterpreting a linear map in the category of `R`-modules -/ -scoped[ModuleCat] notation "↟" f:1024 => ModuleCat.asHom f +scoped[ModuleCat] notation "↟" f:1024 => ModuleCat.ofHom f @[simp 1100] -theorem ModuleCat.asHom_apply {R : Type u} [Ring R] {X Y : Type v} [AddCommGroup X] [Module R X] +theorem ModuleCat.ofHom_apply {R : Type u} [Ring R] {X Y : Type v} [AddCommGroup X] [Module R X] [AddCommGroup Y] [Module R Y] (f : X →ₗ[R] Y) (x : X) : (↟ f) x = f x := rfl -@[deprecated (since := "2024-10-06")] alias ModuleCat.ofHom_apply := ModuleCat.asHom_apply +@[deprecated (since := "2024-10-06")] alias ModuleCat.asHom_apply := ModuleCat.ofHom_apply /-- Reinterpreting a linear map in the category of `R`-modules. -/ def ModuleCat.asHomRight [AddCommGroup X₁] [Module R X₁] {X₂ : ModuleCat.{v} R} : @@ -440,8 +440,8 @@ end ModuleCat @[simp] theorem LinearMap.comp_id_moduleCat {R} [Ring R] {G : ModuleCat.{u} R} {H : Type u} [AddCommGroup H] [Module R H] (f : G →ₗ[R] H) : f.comp (𝟙 G) = f := - Category.id_comp (ModuleCat.asHom f) + Category.id_comp (ModuleCat.ofHom f) @[simp] theorem LinearMap.id_moduleCat_comp {R} [Ring R] {G : Type u} [AddCommGroup G] [Module R G] {H : ModuleCat.{u} R} (f : G →ₗ[R] H) : LinearMap.comp (𝟙 H) f = f := - Category.comp_id (ModuleCat.asHom f) + Category.comp_id (ModuleCat.ofHom f) diff --git a/Mathlib/Algebra/Category/ModuleCat/Biproducts.lean b/Mathlib/Algebra/Category/ModuleCat/Biproducts.lean index b1862b543607c..3c026878abc0c 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Biproducts.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Biproducts.lean @@ -144,8 +144,8 @@ of modules. -/ noncomputable def lequivProdOfRightSplitExact {f : B →ₗ[R] M} (hj : Function.Injective j) (exac : LinearMap.range j = LinearMap.ker g) (h : g.comp f = LinearMap.id) : (A × B) ≃ₗ[R] M := ((ShortComplex.Splitting.ofExactOfSection _ - (ShortComplex.Exact.moduleCat_of_range_eq_ker (ModuleCat.asHom j) - (ModuleCat.asHom g) exac) (asHom f) h + (ShortComplex.Exact.moduleCat_of_range_eq_ker (ModuleCat.ofHom j) + (ModuleCat.ofHom g) exac) (ofHom f) h (by simpa only [ModuleCat.mono_iff_injective])).isoBinaryBiproduct ≪≫ biprodIsoProd _ _ ).symm.toLinearEquiv @@ -154,8 +154,8 @@ of modules. -/ noncomputable def lequivProdOfLeftSplitExact {f : M →ₗ[R] A} (hg : Function.Surjective g) (exac : LinearMap.range j = LinearMap.ker g) (h : f.comp j = LinearMap.id) : (A × B) ≃ₗ[R] M := ((ShortComplex.Splitting.ofExactOfRetraction _ - (ShortComplex.Exact.moduleCat_of_range_eq_ker (ModuleCat.asHom j) - (ModuleCat.asHom g) exac) (ModuleCat.asHom f) h + (ShortComplex.Exact.moduleCat_of_range_eq_ker (ModuleCat.ofHom j) + (ModuleCat.ofHom g) exac) (ModuleCat.ofHom f) h (by simpa only [ModuleCat.epi_iff_surjective] using hg)).isoBinaryBiproduct ≪≫ biprodIsoProd _ _).symm.toLinearEquiv diff --git a/Mathlib/Algebra/Category/ModuleCat/Images.lean b/Mathlib/Algebra/Category/ModuleCat/Images.lean index 9a516ef12b9b4..d37543fe17e12 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Images.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Images.lean @@ -97,12 +97,12 @@ noncomputable def imageIsoRange {G H : ModuleCat.{v} R} (f : G ⟶ H) : @[simp, reassoc, elementwise] theorem imageIsoRange_inv_image_ι {G H : ModuleCat.{v} R} (f : G ⟶ H) : - (imageIsoRange f).inv ≫ Limits.image.ι f = ModuleCat.asHom f.range.subtype := + (imageIsoRange f).inv ≫ Limits.image.ι f = ModuleCat.ofHom f.range.subtype := IsImage.isoExt_inv_m _ _ @[simp, reassoc, elementwise] theorem imageIsoRange_hom_subtype {G H : ModuleCat.{v} R} (f : G ⟶ H) : - (imageIsoRange f).hom ≫ ModuleCat.asHom f.range.subtype = Limits.image.ι f := by + (imageIsoRange f).hom ≫ ModuleCat.ofHom f.range.subtype = Limits.image.ι f := by rw [← imageIsoRange_inv_image_ι f, Iso.hom_inv_id_assoc] end ModuleCat diff --git a/Mathlib/Algebra/Category/ModuleCat/Injective.lean b/Mathlib/Algebra/Category/ModuleCat/Injective.lean index f579799501b6a..972d5bbec68c8 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Injective.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Injective.lean @@ -27,8 +27,8 @@ theorem injective_module_of_injective_object [inj : CategoryTheory.Injective <| ModuleCat.of R M] : Module.Injective R M where out X Y _ _ _ _ f hf g := by - have : CategoryTheory.Mono (ModuleCat.asHom f) := (ModuleCat.mono_iff_injective _).mpr hf - obtain ⟨l, rfl⟩ := inj.factors (ModuleCat.asHom g) (ModuleCat.asHom f) + have : CategoryTheory.Mono (ModuleCat.ofHom f) := (ModuleCat.mono_iff_injective _).mpr hf + obtain ⟨l, rfl⟩ := inj.factors (ModuleCat.ofHom g) (ModuleCat.ofHom f) exact ⟨l, fun _ ↦ rfl⟩ theorem injective_iff_injective_object : diff --git a/Mathlib/Algebra/Category/ModuleCat/Kernels.lean b/Mathlib/Algebra/Category/ModuleCat/Kernels.lean index 0a06c74a2fe68..b74fb6f74650b 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Kernels.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Kernels.lean @@ -26,7 +26,7 @@ variable {M N : ModuleCat.{v} R} (f : M ⟶ N) /-- The kernel cone induced by the concrete kernel. -/ def kernelCone : KernelFork f := -- Porting note: previously proven by tidy - KernelFork.ofι (asHom f.ker.subtype) <| by ext x; cases x; assumption + KernelFork.ofι (ofHom f.ker.subtype) <| by ext x; cases x; assumption /-- The kernel of a linear map is a kernel in the categorical sense. -/ def kernelIsLimit : IsLimit (kernelCone f) := @@ -44,7 +44,7 @@ def kernelIsLimit : IsLimit (kernelCone f) := /-- The cokernel cocone induced by the projection onto the quotient. -/ def cokernelCocone : CokernelCofork f := - CokernelCofork.ofπ (asHom f.range.mkQ) <| LinearMap.range_mkQ_comp _ + CokernelCofork.ofπ (ofHom f.range.mkQ) <| LinearMap.range_mkQ_comp _ /-- The projection onto the quotient is a cokernel in the categorical sense. -/ def cokernelIsColimit : IsColimit (cokernelCocone f) := @@ -53,10 +53,10 @@ def cokernelIsColimit : IsColimit (cokernelCocone f) := f.range.liftQ (Cofork.π s) <| LinearMap.range_le_ker_iff.2 <| CokernelCofork.condition s) (fun s => f.range.liftQ_mkQ (Cofork.π s) _) fun s m h => by -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11036): broken dot notation - haveI : Epi (asHom (LinearMap.range f).mkQ) := + haveI : Epi (ofHom (LinearMap.range f).mkQ) := (epi_iff_range_eq_top _).mpr (Submodule.range_mkQ _) -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11036): broken dot notation - apply (cancel_epi (asHom (LinearMap.range f).mkQ)).1 + apply (cancel_epi (ofHom (LinearMap.range f).mkQ)).1 convert h -- Porting note: no longer necessary -- exact Submodule.liftQ_mkQ _ _ _ diff --git a/Mathlib/Algebra/Homology/ShortComplex/ModuleCat.lean b/Mathlib/Algebra/Homology/ShortComplex/ModuleCat.lean index 59034f288455b..7523eadd214dd 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/ModuleCat.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/ModuleCat.lean @@ -33,7 +33,7 @@ linear maps `f` and `g` and the vanishing of their composition. -/ def moduleCatMk {X₁ X₂ X₃ : Type v} [AddCommGroup X₁] [AddCommGroup X₂] [AddCommGroup X₃] [Module R X₁] [Module R X₂] [Module R X₃] (f : X₁ →ₗ[R] X₂) (g : X₂ →ₗ[R] X₃) (hfg : g.comp f = 0) : ShortComplex (ModuleCat.{v} R) := - ShortComplex.mk (ModuleCat.asHom f) (ModuleCat.asHom g) hfg + ShortComplex.mk (ModuleCat.ofHom f) (ModuleCat.ofHom g) hfg variable (S : ShortComplex (ModuleCat.{v} R)) @@ -138,7 +138,7 @@ def moduleCatLeftHomologyData : S.LeftHomologyData where erw [Submodule.Quotient.mk_eq_zero] rw [LinearMap.mem_range] apply exists_apply_eq_apply - hπ := ModuleCat.cokernelIsColimit (ModuleCat.asHom S.moduleCatToCycles) + hπ := ModuleCat.cokernelIsColimit (ModuleCat.ofHom S.moduleCatToCycles) @[simp] lemma moduleCatLeftHomologyData_f' : diff --git a/Mathlib/CategoryTheory/Linear/Yoneda.lean b/Mathlib/CategoryTheory/Linear/Yoneda.lean index ee468da7d7708..4836d3c76a695 100644 --- a/Mathlib/CategoryTheory/Linear/Yoneda.lean +++ b/Mathlib/CategoryTheory/Linear/Yoneda.lean @@ -28,7 +28,7 @@ namespace CategoryTheory variable (R : Type w) [Ring R] {C : Type u} [Category.{v} C] [Preadditive C] [Linear R C] variable (C) --- Porting note: inserted specific `ModuleCat.asHom` in the definition of `linearYoneda` +-- Porting note: inserted specific `ModuleCat.ofHom` in the definition of `linearYoneda` -- and similarly in `linearCoyoneda`, otherwise many simp lemmas are not triggered automatically. -- Eventually, doing so allows more proofs to be automatic! /-- The Yoneda embedding for `R`-linear categories `C`, @@ -38,9 +38,9 @@ with value on `Y : Cᵒᵖ` given by `ModuleCat.of R (unop Y ⟶ X)`. -/ def linearYoneda : C ⥤ Cᵒᵖ ⥤ ModuleCat R where obj X := { obj := fun Y => ModuleCat.of R (unop Y ⟶ X) - map := fun f => ModuleCat.asHom (Linear.leftComp R _ f.unop) } + map := fun f => ModuleCat.ofHom (Linear.leftComp R _ f.unop) } map {X₁ X₂} f := - { app := fun Y => @ModuleCat.asHom R _ (Y.unop ⟶ X₁) (Y.unop ⟶ X₂) _ _ _ _ + { app := fun Y => @ModuleCat.ofHom R _ (Y.unop ⟶ X₁) (Y.unop ⟶ X₂) _ _ _ _ (Linear.rightComp R _ f) } /-- The Yoneda embedding for `R`-linear categories `C`, @@ -50,9 +50,9 @@ with value on `X : C` given by `ModuleCat.of R (unop Y ⟶ X)`. -/ def linearCoyoneda : Cᵒᵖ ⥤ C ⥤ ModuleCat R where obj Y := { obj := fun X => ModuleCat.of R (unop Y ⟶ X) - map := fun f => ModuleCat.asHom (Linear.rightComp R _ f) } + map := fun f => ModuleCat.ofHom (Linear.rightComp R _ f) } map {Y₁ Y₂} f := - { app := fun X => @ModuleCat.asHom R _ (unop Y₁ ⟶ X) (unop Y₂ ⟶ X) _ _ _ _ + { app := fun X => @ModuleCat.ofHom R _ (unop Y₁ ⟶ X) (unop Y₂ ⟶ X) _ _ _ _ (Linear.leftComp _ _ f.unop) } instance linearYoneda_obj_additive (X : C) : ((linearYoneda R C).obj X).Additive where diff --git a/Mathlib/CategoryTheory/Preadditive/Yoneda/Basic.lean b/Mathlib/CategoryTheory/Preadditive/Yoneda/Basic.lean index 812bfc902cea1..beedbdb8d0a67 100644 --- a/Mathlib/CategoryTheory/Preadditive/Yoneda/Basic.lean +++ b/Mathlib/CategoryTheory/Preadditive/Yoneda/Basic.lean @@ -39,7 +39,7 @@ object `X` to the `End Y`-module of morphisms `X ⟶ Y`. @[simps] def preadditiveYonedaObj (Y : C) : Cᵒᵖ ⥤ ModuleCat.{v} (End Y) where obj X := ModuleCat.of _ (X.unop ⟶ Y) - map f := ModuleCat.asHom + map f := ModuleCat.ofHom { toFun := fun g => f.unop ≫ g map_add' := fun _ _ => comp_add _ _ _ _ _ _ map_smul' := fun _ _ => Eq.symm <| Category.assoc _ _ _ } @@ -66,7 +66,7 @@ object `Y` to the `End X`-module of morphisms `X ⟶ Y`. @[simps] def preadditiveCoyonedaObj (X : Cᵒᵖ) : C ⥤ ModuleCat.{v} (End X) where obj Y := ModuleCat.of _ (unop X ⟶ Y) - map f := ModuleCat.asHom + map f := ModuleCat.ofHom { toFun := fun g => g ≫ f map_add' := fun _ _ => add_comp _ _ _ _ _ _ map_smul' := fun _ _ => Category.assoc _ _ _ } diff --git a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean index 24d8d8f5bd37c..6f7158e455f66 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean @@ -201,9 +201,9 @@ theorem dOne_comp_dZero : dOne A ∘ₗ dZero A = 0 := by rfl theorem dTwo_comp_dOne : dTwo A ∘ₗ dOne A = 0 := by - show ModuleCat.asHom (dOne A) ≫ ModuleCat.asHom (dTwo A) = _ - have h1 : _ ≫ ModuleCat.asHom (dOne A) = _ ≫ _ := congr_arg ModuleCat.asHom (dOne_comp_eq A) - have h2 : _ ≫ ModuleCat.asHom (dTwo A) = _ ≫ _ := congr_arg ModuleCat.asHom (dTwo_comp_eq A) + show ModuleCat.ofHom (dOne A) ≫ ModuleCat.ofHom (dTwo A) = _ + have h1 : _ ≫ ModuleCat.ofHom (dOne A) = _ ≫ _ := congr_arg ModuleCat.ofHom (dOne_comp_eq A) + have h2 : _ ≫ ModuleCat.ofHom (dTwo A) = _ ≫ _ := congr_arg ModuleCat.ofHom (dTwo_comp_eq A) simp only [← LinearEquiv.toModuleIso_hom] at h1 h2 simp only [(Iso.eq_inv_comp _).2 h2, (Iso.eq_inv_comp _).2 h1, Category.assoc, Iso.hom_inv_id_assoc, HomologicalComplex.d_comp_d_assoc, zero_comp, comp_zero] @@ -749,7 +749,7 @@ lemma shortComplexH0_exact : (shortComplexH0 A).Exact := by `(inhomogeneousCochains A).d 0 1` of the complex of inhomogeneous cochains of `A`. -/ @[simps! hom_left hom_right inv_left inv_right] def dZeroArrowIso : Arrow.mk ((inhomogeneousCochains A).d 0 1) ≅ - Arrow.mk (ModuleCat.asHom (dZero A)) := + Arrow.mk (ModuleCat.ofHom (dZero A)) := Arrow.isoMk (zeroCochainsLequiv A).toModuleIso (oneCochainsLequiv A).toModuleIso (dZero_comp_eq A) @@ -797,7 +797,7 @@ def isoOneCocycles : cocycles A 1 ≅ ModuleCat.of k (oneCocycles A) := cyclesMapIso (shortComplexH1Iso A) ≪≫ (shortComplexH1 A).moduleCatCyclesIso lemma isoOneCocycles_hom_comp_subtype : - (isoOneCocycles A).hom ≫ ModuleCat.asHom (oneCocycles A).subtype = + (isoOneCocycles A).hom ≫ ModuleCat.ofHom (oneCocycles A).subtype = iCocycles A 1 ≫ (oneCochainsLequiv A).toModuleIso.hom := by dsimp [isoOneCocycles] rw [Category.assoc, Category.assoc] @@ -807,7 +807,7 @@ lemma isoOneCocycles_hom_comp_subtype : lemma toCocycles_comp_isoOneCocycles_hom : toCocycles A 0 1 ≫ (isoOneCocycles A).hom = (zeroCochainsLequiv A).toModuleIso.hom ≫ - ModuleCat.asHom (shortComplexH1 A).moduleCatToCycles := by + ModuleCat.ofHom (shortComplexH1 A).moduleCatToCycles := by simp [isoOneCocycles] rfl @@ -845,7 +845,7 @@ def isoTwoCocycles : cocycles A 2 ≅ ModuleCat.of k (twoCocycles A) := cyclesMapIso (shortComplexH2Iso A) ≪≫ (shortComplexH2 A).moduleCatCyclesIso lemma isoTwoCocycles_hom_comp_subtype : - (isoTwoCocycles A).hom ≫ ModuleCat.asHom (twoCocycles A).subtype = + (isoTwoCocycles A).hom ≫ ModuleCat.ofHom (twoCocycles A).subtype = iCocycles A 2 ≫ (twoCochainsLequiv A).toModuleIso.hom := by dsimp [isoTwoCocycles] rw [Category.assoc, Category.assoc] @@ -855,7 +855,7 @@ lemma isoTwoCocycles_hom_comp_subtype : lemma toCocycles_comp_isoTwoCocycles_hom : toCocycles A 1 2 ≫ (isoTwoCocycles A).hom = (oneCochainsLequiv A).toModuleIso.hom ≫ - ModuleCat.asHom (shortComplexH2 A).moduleCatToCycles := by + ModuleCat.ofHom (shortComplexH2 A).moduleCatToCycles := by simp [isoTwoCocycles] rfl diff --git a/Mathlib/RepresentationTheory/Rep.lean b/Mathlib/RepresentationTheory/Rep.lean index 0a14328b88a96..e445fe7584092 100644 --- a/Mathlib/RepresentationTheory/Rep.lean +++ b/Mathlib/RepresentationTheory/Rep.lean @@ -332,7 +332,7 @@ variable [Group G] (A B C : Rep k G) protected def ihom (A : Rep k G) : Rep k G ⥤ Rep k G where obj B := Rep.of (Representation.linHom A.ρ B.ρ) map := fun {X} {Y} f => - { hom := ModuleCat.asHom (LinearMap.llcomp k _ _ _ f.hom) + { hom := ModuleCat.ofHom (LinearMap.llcomp k _ _ _ f.hom) comm := fun g => LinearMap.ext fun x => LinearMap.ext fun y => by show f.hom (X.ρ g _) = _ simp only [hom_comm_apply]; rfl } diff --git a/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean b/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean index 1c4dfed5cf15a..afc24e101812b 100644 --- a/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean +++ b/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean @@ -241,17 +241,17 @@ private instance : AddCommGroup (AdicCompletion I R ⊗[R] (LinearMap.ker f)) := private def firstRow : ComposableArrows (ModuleCat (AdicCompletion I R)) 4 := ComposableArrows.mk₄ - (ModuleCat.asHom <| lTensorKerIncl I M f) - (ModuleCat.asHom <| lTensorf I M f) - (ModuleCat.asHom (0 : AdicCompletion I R ⊗[R] M →ₗ[AdicCompletion I R] PUnit)) - (ModuleCat.asHom (0 : _ →ₗ[AdicCompletion I R] PUnit)) + (ModuleCat.ofHom <| lTensorKerIncl I M f) + (ModuleCat.ofHom <| lTensorf I M f) + (ModuleCat.ofHom (0 : AdicCompletion I R ⊗[R] M →ₗ[AdicCompletion I R] PUnit)) + (ModuleCat.ofHom (0 : _ →ₗ[AdicCompletion I R] PUnit)) private def secondRow : ComposableArrows (ModuleCat (AdicCompletion I R)) 4 := ComposableArrows.mk₄ - (ModuleCat.asHom (map I <| (LinearMap.ker f).subtype)) - (ModuleCat.asHom (map I f)) - (ModuleCat.asHom (0 : _ →ₗ[AdicCompletion I R] PUnit)) - (ModuleCat.asHom (0 : _ →ₗ[AdicCompletion I R] PUnit)) + (ModuleCat.ofHom (map I <| (LinearMap.ker f).subtype)) + (ModuleCat.ofHom (map I f)) + (ModuleCat.ofHom (0 : _ →ₗ[AdicCompletion I R] PUnit)) + (ModuleCat.ofHom (0 : _ →ₗ[AdicCompletion I R] PUnit)) include hf @@ -282,25 +282,25 @@ private lemma secondRow_exact [Fintype ι] [IsNoetherianRing R] : (secondRow I M /- The compatible vertical maps between the first and the second row. -/ private def firstRowToSecondRow : firstRow I M f ⟶ secondRow I M f := ComposableArrows.homMk₄ - (ModuleCat.asHom (ofTensorProduct I (LinearMap.ker f))) - (ModuleCat.asHom (ofTensorProduct I (ι → R))) - (ModuleCat.asHom (ofTensorProduct I M)) - (ModuleCat.asHom 0) - (ModuleCat.asHom 0) + (ModuleCat.ofHom (ofTensorProduct I (LinearMap.ker f))) + (ModuleCat.ofHom (ofTensorProduct I (ι → R))) + (ModuleCat.ofHom (ofTensorProduct I M)) + (ModuleCat.ofHom 0) + (ModuleCat.ofHom 0) (ofTensorProduct_naturality I <| (LinearMap.ker f).subtype).symm (ofTensorProduct_naturality I f).symm rfl rfl private lemma ofTensorProduct_iso [Fintype ι] [IsNoetherianRing R] : - IsIso (ModuleCat.asHom (ofTensorProduct I M)) := by + IsIso (ModuleCat.ofHom (ofTensorProduct I M)) := by refine Abelian.isIso_of_epi_of_isIso_of_isIso_of_mono (firstRow_exact I M f hf) (secondRow_exact I M f hf) (firstRowToSecondRow I M f) ?_ ?_ ?_ ?_ · apply ConcreteCategory.epi_of_surjective exact ofTensorProduct_surjective_of_finite I (LinearMap.ker f) · apply (ConcreteCategory.isIso_iff_bijective _).mpr exact ofTensorProduct_bijective_of_pi_of_fintype I ι - · show IsIso (ModuleCat.asHom 0) + · show IsIso (ModuleCat.ofHom 0) apply Limits.isIso_of_isTerminal <;> exact Limits.IsZero.isTerminal (ModuleCat.isZero_of_subsingleton _) · apply ConcreteCategory.mono_of_injective @@ -310,9 +310,9 @@ private lemma ofTensorProduct_iso [Fintype ι] [IsNoetherianRing R] : private lemma ofTensorProduct_bijective_of_map_from_fin [Fintype ι] [IsNoetherianRing R] : Function.Bijective (ofTensorProduct I M) := by - have : IsIso (ModuleCat.asHom (ofTensorProduct I M)) := + have : IsIso (ModuleCat.ofHom (ofTensorProduct I M)) := ofTensorProduct_iso I M f hf - exact ConcreteCategory.bijective_of_isIso (ModuleCat.asHom (ofTensorProduct I M)) + exact ConcreteCategory.bijective_of_isIso (ModuleCat.ofHom (ofTensorProduct I M)) end diff --git a/Mathlib/RingTheory/Coalgebra/TensorProduct.lean b/Mathlib/RingTheory/Coalgebra/TensorProduct.lean index c7ac28b39f0ee..9efad5fea6593 100644 --- a/Mathlib/RingTheory/Coalgebra/TensorProduct.lean +++ b/Mathlib/RingTheory/Coalgebra/TensorProduct.lean @@ -58,7 +58,7 @@ noncomputable instance TensorProduct.instCoalgebra : Coalgebra R (M ⊗[R] N) := rw [CoalgebraCat.ofComonObjCoalgebraStruct_comul] simp [-Mon_.monMonoidalStruct_tensorObj_X, ModuleCat.MonoidalCategory.instMonoidalCategoryStruct_tensorHom, - ModuleCat.comp_def, ModuleCat.of, ModuleCat.asHom, + ModuleCat.comp_def, ModuleCat.of, ModuleCat.ofHom, ModuleCat.MonoidalCategory.tensorμ_eq_tensorTensorTensorComm] } end diff --git a/Mathlib/RingTheory/Flat/CategoryTheory.lean b/Mathlib/RingTheory/Flat/CategoryTheory.lean index ea6d1442dad38..ea7bc816d2f57 100644 --- a/Mathlib/RingTheory/Flat/CategoryTheory.lean +++ b/Mathlib/RingTheory/Flat/CategoryTheory.lean @@ -52,7 +52,7 @@ lemma iff_lTensor_preserves_shortComplex_exact : ⟨fun _ _ ↦ lTensor_shortComplex_exact _ _, fun H ↦ iff_lTensor_exact.2 <| fun _ _ _ _ _ _ _ _ _ f g h ↦ moduleCat_exact_iff_function_exact _ |>.1 <| - H (.mk (ModuleCat.asHom f) (ModuleCat.asHom g) + H (.mk (ModuleCat.ofHom f) (ModuleCat.ofHom g) (DFunLike.ext _ _ h.apply_apply_eq_zero)) (moduleCat_exact_iff_function_exact _ |>.2 h)⟩ @@ -62,7 +62,7 @@ lemma iff_rTensor_preserves_shortComplex_exact : ⟨fun _ _ ↦ rTensor_shortComplex_exact _ _, fun H ↦ iff_rTensor_exact.2 <| fun _ _ _ _ _ _ _ _ _ f g h ↦ moduleCat_exact_iff_function_exact _ |>.1 <| - H (.mk (ModuleCat.asHom f) (ModuleCat.asHom g) + H (.mk (ModuleCat.ofHom f) (ModuleCat.ofHom g) (DFunLike.ext _ _ h.apply_apply_eq_zero)) (moduleCat_exact_iff_function_exact _ |>.2 h)⟩ diff --git a/Mathlib/Topology/Category/Profinite/Nobeling.lean b/Mathlib/Topology/Category/Profinite/Nobeling.lean index 42749f3df74e4..f421a36182178 100644 --- a/Mathlib/Topology/Category/Profinite/Nobeling.lean +++ b/Mathlib/Topology/Category/Profinite/Nobeling.lean @@ -1354,13 +1354,13 @@ theorem CC_exact {f : LocallyConstant C ℤ} (hf : Linear_CC' C hsC ho f = 0) : exact C1_projOrd C hsC ho hx₁ variable (o) in -theorem succ_mono : CategoryTheory.Mono (ModuleCat.asHom (πs C o)) := by +theorem succ_mono : CategoryTheory.Mono (ModuleCat.ofHom (πs C o)) := by rw [ModuleCat.mono_iff_injective] exact injective_πs _ _ include hC in theorem succ_exact : - (ShortComplex.mk (ModuleCat.asHom (πs C o)) (ModuleCat.asHom (Linear_CC' C hsC ho)) + (ShortComplex.mk (ModuleCat.ofHom (πs C o)) (ModuleCat.ofHom (Linear_CC' C hsC ho)) (by ext; apply CC_comp_zero)).Exact := by rw [ShortComplex.moduleCat_exact_iff] intro f @@ -1478,7 +1478,7 @@ theorem span_sum : Set.range (eval C) = Set.range (Sum.elim theorem square_commutes : SumEval C ho ∘ Sum.inl = - ModuleCat.asHom (πs C o) ∘ eval (π C (ord I · < o)) := by + ModuleCat.ofHom (πs C o) ∘ eval (π C (ord I · < o)) := by ext l dsimp [SumEval] rw [← Products.eval_πs C (Products.prop_of_isGood _ _ l.prop)] @@ -1644,7 +1644,7 @@ theorem maxTail_isGood (l : MaxProducts C ho) rfl have hse := succ_exact C hC hsC ho rw [ShortComplex.moduleCat_exact_iff_range_eq_ker] at hse - dsimp [ModuleCat.asHom] at hse + dsimp [ModuleCat.ofHom] at hse -- Rewrite `this` using exact sequence manipulations to conclude that a term is in the range of -- the linear map `πs`: @@ -1701,8 +1701,8 @@ include hC in theorem linearIndependent_comp_of_eval (h₁ : ⊤ ≤ Submodule.span ℤ (Set.range (eval (π C (ord I · < o))))) : LinearIndependent ℤ (eval (C' C ho)) → - LinearIndependent ℤ (ModuleCat.asHom (Linear_CC' C hsC ho) ∘ SumEval C ho ∘ Sum.inr) := by - dsimp [SumEval, ModuleCat.asHom] + LinearIndependent ℤ (ModuleCat.ofHom (Linear_CC' C hsC ho) ∘ SumEval C ho ∘ Sum.inr) := by + dsimp [SumEval, ModuleCat.ofHom] erw [max_eq_eval_unapply C hsC ho] intro h let f := MaxToGood C hC hsC ho h₁ From a8931809d8cef5ce0c893acc2f907ecef61e8ac8 Mon Sep 17 00:00:00 2001 From: Scott Carnahan <128885296+ScottCarnahan@users.noreply.github.com> Date: Tue, 3 Dec 2024 14:52:37 +0000 Subject: [PATCH 495/829] feat (RingTheory/HahnSeries/Summable) add pointwise product operation (#18794) This PR defines a pointwise product operation on summable families of Hahn series, and shows that the formal sum of the pointwise product is equal to the product of formal sums. This is preparation for defining algebra homomorphisms from power series rings. --- Mathlib/RingTheory/HahnSeries/Summable.lean | 77 ++++++++++++++++++--- 1 file changed, 66 insertions(+), 11 deletions(-) diff --git a/Mathlib/RingTheory/HahnSeries/Summable.lean b/Mathlib/RingTheory/HahnSeries/Summable.lean index c9ea82db6087b..41f0ce1b8f513 100644 --- a/Mathlib/RingTheory/HahnSeries/Summable.lean +++ b/Mathlib/RingTheory/HahnSeries/Summable.lean @@ -267,7 +267,22 @@ end AddCommGroup section SMul -variable [PartialOrder Γ] [PartialOrder Γ'] [AddCommMonoid R] [AddCommMonoid V] [SMulWithZero R V] +variable [PartialOrder Γ] [PartialOrder Γ'] [AddCommMonoid V] + +instance [Zero R] [SMulWithZero R V] : SMul R (SummableFamily Γ' V β) := + ⟨fun r t => + { toFun := r • t + isPWO_iUnion_support' := t.isPWO_iUnion_support.mono (Set.iUnion_mono fun i => + Pi.smul_apply r t i ▸ Function.support_const_smul_subset r _) + finite_co_support' := by + intro g + refine (t.finite_co_support g).subset ?_ + intro i hi + simp only [Pi.smul_apply, smul_coeff, ne_eq, Set.mem_setOf_eq] at hi + simp only [Function.mem_support, ne_eq] + exact right_ne_zero_of_smul hi } ⟩ + +variable [AddCommMonoid R] [SMulWithZero R V] theorem smul_support_subset_prod (s : SummableFamily Γ R α) (t : SummableFamily Γ' V β) (gh : Γ × Γ') : @@ -322,13 +337,15 @@ theorem finite_co_support_prod_smul (s : SummableFamily Γ R α) /-- An elementwise scalar multiplication of one summable family on another. -/ @[simps] -def FamilySMul (s : SummableFamily Γ R α) (t : SummableFamily Γ' V β) : +def smul (s : SummableFamily Γ R α) (t : SummableFamily Γ' V β) : (SummableFamily Γ' V (α × β)) where toFun ab := (of R).symm (s (ab.1) • ((of R) (t (ab.2)))) isPWO_iUnion_support' := isPWO_iUnion_support_prod_smul s.isPWO_iUnion_support t.isPWO_iUnion_support finite_co_support' g := finite_co_support_prod_smul s t g +@[deprecated (since := "2024-11-17")] noncomputable alias FamilySMul := smul + theorem sum_vAddAntidiagonal_eq (s : SummableFamily Γ R α) (t : SummableFamily Γ' V β) (g : Γ') (a : α × β) : ∑ x ∈ VAddAntidiagonal (s a.1).isPWO_support' (t a.2).isPWO_support' g, (s a.1).coeff x.1 • @@ -341,12 +358,12 @@ theorem sum_vAddAntidiagonal_eq (s : SummableFamily Γ R α) (t : SummableFamily · exact smul_eq_zero_of_left hs ((t a.2).coeff gh.2) · simp_all -theorem family_smul_coeff {R} {V} [Semiring R] [AddCommMonoid V] [Module R V] +theorem smul_coeff {R} {V} [Semiring R] [AddCommMonoid V] [Module R V] (s : SummableFamily Γ R α) (t : SummableFamily Γ' V β) (g : Γ') : - (FamilySMul s t).hsum.coeff g = ∑ gh ∈ VAddAntidiagonal s.isPWO_iUnion_support + (smul s t).hsum.coeff g = ∑ gh ∈ VAddAntidiagonal s.isPWO_iUnion_support t.isPWO_iUnion_support g, (s.hsum.coeff gh.1) • (t.hsum.coeff gh.2) := by rw [hsum_coeff] - simp only [hsum_coeff_eq_sum, FamilySMul_toFun, HahnModule.smul_coeff, Equiv.symm_apply_apply] + simp only [hsum_coeff_eq_sum, smul_toFun, HahnModule.smul_coeff, Equiv.symm_apply_apply] simp_rw [sum_vAddAntidiagonal_eq, Finset.smul_sum, Finset.sum_smul] rw [← sum_finsum_comm _ _ <| fun gh _ => smul_support_finite s t gh] refine sum_congr rfl fun gh _ => ?_ @@ -358,11 +375,13 @@ theorem family_smul_coeff {R} {V} [Semiring R] [AddCommMonoid V] [Module R V] Set.mem_setOf_eq, Prod.forall, coeff_support, mem_product] exact hsupp ab.1 ab.2 hab -theorem hsum_family_smul {R} {V} [Semiring R] [AddCommMonoid V] [Module R V] +@[deprecated (since := "2024-11-17")] alias family_smul_coeff := smul_coeff + +theorem smul_hsum {R} {V} [Semiring R] [AddCommMonoid V] [Module R V] (s : SummableFamily Γ R α) (t : SummableFamily Γ' V β) : - (FamilySMul s t).hsum = (of R).symm (s.hsum • (of R) (t.hsum)) := by + (smul s t).hsum = (of R).symm (s.hsum • (of R) (t.hsum)) := by ext g - rw [family_smul_coeff s t g, HahnModule.smul_coeff, Equiv.symm_apply_apply] + rw [smul_coeff s t g, HahnModule.smul_coeff, Equiv.symm_apply_apply] refine Eq.symm (sum_of_injOn (fun a ↦ a) (fun _ _ _ _ h ↦ h) (fun _ hgh => ?_) (fun gh _ hgh => ?_) fun _ _ => by simp) · simp_all only [mem_coe, mem_vaddAntidiagonal, mem_support, ne_eq, Set.mem_iUnion, and_true] @@ -377,11 +396,13 @@ theorem hsum_family_smul {R} {V} [Semiring R] [AddCommMonoid V] [Module R V] · exact smul_eq_zero_of_left h (t.hsum.coeff gh.2) · simp_all +@[deprecated (since := "2024-11-17")] alias hsum_family_smul := smul_hsum + instance [AddCommMonoid R] [SMulWithZero R V] : SMul (HahnSeries Γ R) (SummableFamily Γ' V β) where - smul x t := Equiv (Equiv.punitProd β) <| FamilySMul (single x) t + smul x t := Equiv (Equiv.punitProd β) <| smul (single x) t theorem smul_eq {x : HahnSeries Γ R} {t : SummableFamily Γ' V β} : - x • t = Equiv (Equiv.punitProd β) (FamilySMul (single x) t) := + x • t = Equiv (Equiv.punitProd β) (smul (single x) t) := rfl @[simp] @@ -393,7 +414,7 @@ theorem smul_apply {x : HahnSeries Γ R} {s : SummableFamily Γ' V α} {a : α} theorem hsum_smul_module {R} {V} [Semiring R] [AddCommMonoid V] [Module R V] {x : HahnSeries Γ R} {s : SummableFamily Γ' V α} : (x • s).hsum = (of R).symm (x • of R s.hsum) := by - rw [smul_eq, hsum_equiv, hsum_family_smul, hsum_single] + rw [smul_eq, hsum_equiv, smul_hsum, hsum_single] end SMul @@ -427,6 +448,40 @@ theorem hsum_sub {R : Type*} [Ring R] {s t : SummableFamily Γ R α} : (s - t).hsum = s.hsum - t.hsum := by rw [← lsum_apply, LinearMap.map_sub, lsum_apply, lsum_apply] +theorem isPWO_iUnion_support_prod_mul {s : α → HahnSeries Γ R} {t : β → HahnSeries Γ R} + (hs : (⋃ a, (s a).support).IsPWO) (ht : (⋃ b, (t b).support).IsPWO) : + (⋃ (a : α × β), ((fun a ↦ ((s a.1) * (t a.2))) a).support).IsPWO := + isPWO_iUnion_support_prod_smul hs ht + +theorem finite_co_support_prod_mul (s : SummableFamily Γ R α) + (t : SummableFamily Γ R β) (g : Γ) : + Finite {(a : α × β) | ((fun (a : α × β) ↦ (s a.1 * t a.2)) a).coeff g ≠ 0} := + finite_co_support_prod_smul s t g + +/-- A summable family given by pointwise multiplication of a pair of summable families. -/ +@[simps] +def mul (s : SummableFamily Γ R α) (t : SummableFamily Γ R β) : + (SummableFamily Γ R (α × β)) where + toFun a := s (a.1) * t (a.2) + isPWO_iUnion_support' := + isPWO_iUnion_support_prod_mul s.isPWO_iUnion_support t.isPWO_iUnion_support + finite_co_support' g := finite_co_support_prod_mul s t g + +theorem mul_eq_smul {β : Type*} (s : SummableFamily Γ R α) (t : SummableFamily Γ R β) : + mul s t = smul s t := + rfl + +theorem mul_coeff {β : Type*} (s : SummableFamily Γ R α) (t : SummableFamily Γ R β) (g : Γ) : + (mul s t).hsum.coeff g = ∑ gh ∈ addAntidiagonal s.isPWO_iUnion_support + t.isPWO_iUnion_support g, (s.hsum.coeff gh.1) * (t.hsum.coeff gh.2) := by + simp_rw [← smul_eq_mul, mul_eq_smul] + exact smul_coeff s t g + +theorem hsum_mul {β : Type*} (s : SummableFamily Γ R α) (t : SummableFamily Γ R β) : + (mul s t).hsum = s.hsum * t.hsum := by + rw [← smul_eq_mul, mul_eq_smul] + exact smul_hsum s t + end Semiring section OfFinsupp From c5be7923a1fa0766edbfb303eeaaa683886b23fc Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9s=20Goens?= Date: Tue, 3 Dec 2024 15:10:08 +0000 Subject: [PATCH 496/829] =?UTF-8?q?refactor(Quot):=20use=20dependent=20arr?= =?UTF-8?q?ows=20in=20congruence=20arguments=20for=20=20`map`=20and=20`map?= =?UTF-8?q?=E2=82=82`=20=20(#15127)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The `map` and `map₂` functions for `Quot` and `Quotient` used a Lean 3 version of congruence arguments that avoids dependent arrows. This PR changes those arguments to have the more recognizable formulation of congruence. Co-authored-by: Andrés Goens Co-authored-by: Andrés Goens --- Mathlib/Algebra/GCDMonoid/Basic.lean | 4 +-- Mathlib/Data/Quot.lean | 26 +++++++------------ Mathlib/Data/Setoid/Basic.lean | 2 +- Mathlib/GroupTheory/Congruence/Defs.lean | 4 +-- .../Projectivization/Constructions.lean | 2 +- Mathlib/Order/Filter/Germ/Basic.lean | 2 +- .../HomogeneousLocalization.lean | 4 +-- .../Algebra/SeparationQuotient/Basic.lean | 4 +-- 8 files changed, 21 insertions(+), 27 deletions(-) diff --git a/Mathlib/Algebra/GCDMonoid/Basic.lean b/Mathlib/Algebra/GCDMonoid/Basic.lean index 917885981d942..b6ff248ab4b41 100644 --- a/Mathlib/Algebra/GCDMonoid/Basic.lean +++ b/Mathlib/Algebra/GCDMonoid/Basic.lean @@ -1315,8 +1315,8 @@ namespace Associates variable [CancelCommMonoidWithZero α] [GCDMonoid α] instance instGCDMonoid : GCDMonoid (Associates α) where - gcd := Quotient.map₂' gcd fun _ _ (ha : Associated _ _) _ _ (hb : Associated _ _) => ha.gcd hb - lcm := Quotient.map₂' lcm fun _ _ (ha : Associated _ _) _ _ (hb : Associated _ _) => ha.lcm hb + gcd := Quotient.map₂ gcd fun _ _ (ha : Associated _ _) _ _ (hb : Associated _ _) => ha.gcd hb + lcm := Quotient.map₂ lcm fun _ _ (ha : Associated _ _) _ _ (hb : Associated _ _) => ha.lcm hb gcd_dvd_left := by rintro ⟨a⟩ ⟨b⟩; exact mk_le_mk_of_dvd (gcd_dvd_left _ _) gcd_dvd_right := by rintro ⟨a⟩ ⟨b⟩; exact mk_le_mk_of_dvd (gcd_dvd_right _ _) dvd_gcd := by diff --git a/Mathlib/Data/Quot.lean b/Mathlib/Data/Quot.lean index 9508faa6a1ad6..ccad963fe4394 100644 --- a/Mathlib/Data/Quot.lean +++ b/Mathlib/Data/Quot.lean @@ -79,8 +79,8 @@ protected def hrecOn₂ (qa : Quot ra) (qb : Quot rb) (f : ∀ a b, φ ⟦a⟧ /-- Map a function `f : α → β` such that `ra x y` implies `rb (f x) (f y)` to a map `Quot ra → Quot rb`. -/ -protected def map (f : α → β) (h : (ra ⇒ rb) f f) : Quot ra → Quot rb := - (Quot.lift fun x ↦ ⟦f x⟧) fun x y (h₁ : ra x y) ↦ Quot.sound <| h h₁ +protected def map (f : α → β) (h : ∀ ⦃a b : α⦄, ra a b → rb (f a) (f b)) : Quot ra → Quot rb := + Quot.lift (fun x => Quot.mk rb (f x)) fun _ _ hra ↦ Quot.sound <| h hra /-- If `ra` is a subrelation of `ra'`, then we have a natural map `Quot ra → Quot ra'`. -/ protected def mapRight {ra' : α → α → Prop} (h : ∀ a₁ a₂, ra a₁ a₂ → ra' a₁ a₂) : @@ -227,11 +227,11 @@ protected def hrecOn₂ (qa : Quotient sa) (qb : Quotient sb) (f : ∀ a b, φ /-- Map a function `f : α → β` that sends equivalent elements to equivalent elements to a function `Quotient sa → Quotient sb`. Useful to define unary operations on quotients. -/ -protected def map (f : α → β) (h : ((· ≈ ·) ⇒ (· ≈ ·)) f f) : Quotient sa → Quotient sb := +protected def map (f : α → β) (h : ∀ ⦃a b : α⦄, a ≈ b → f a ≈ f b) : Quotient sa → Quotient sb := Quot.map f h @[simp] -theorem map_mk (f : α → β) (h : ((· ≈ ·) ⇒ (· ≈ ·)) f f) (x : α) : +theorem map_mk (f : α → β) (h) (x : α) : Quotient.map f h (⟦x⟧ : Quotient sa) = (⟦f x⟧ : Quotient sb) := rfl @@ -240,12 +240,13 @@ variable {γ : Sort*} {sc : Setoid γ} /-- Map a function `f : α → β → γ` that sends equivalent elements to equivalent elements to a function `f : Quotient sa → Quotient sb → Quotient sc`. Useful to define binary operations on quotients. -/ -protected def map₂ (f : α → β → γ) (h : ((· ≈ ·) ⇒ (· ≈ ·) ⇒ (· ≈ ·)) f f) : +protected def map₂ (f : α → β → γ) + (h : ∀ ⦃a₁ a₂⦄, a₁ ≈ a₂ → ∀ ⦃b₁ b₂⦄, b₁ ≈ b₂ → f a₁ b₁ ≈ f a₂ b₂) : Quotient sa → Quotient sb → Quotient sc := Quotient.lift₂ (fun x y ↦ ⟦f x y⟧) fun _ _ _ _ h₁ h₂ ↦ Quot.sound <| h h₁ h₂ @[simp] -theorem map₂_mk (f : α → β → γ) (h : ((· ≈ ·) ⇒ (· ≈ ·) ⇒ (· ≈ ·)) f f) (x : α) (y : β) : +theorem map₂_mk (f : α → β → γ) (h) (x : α) (y : β) : Quotient.map₂ f h (⟦x⟧ : Quotient sa) (⟦y⟧ : Quotient sb) = (⟦f x y⟧ : Quotient sc) := rfl @@ -693,7 +694,8 @@ theorem hrecOn₂'_mk'' {φ : Quotient s₁ → Quotient s₂ → Sort*} /-- Map a function `f : α → β` that sends equivalent elements to equivalent elements to a function `Quotient sa → Quotient sb`. Useful to define unary operations on quotients. -/ -protected def map' (f : α → β) (h : (s₁.r ⇒ s₂.r) f f) : Quotient s₁ → Quotient s₂ := +protected def map' (f : α → β) (h : ∀ a b, s₁.r a b → s₂.r (f a) (f b)) : + Quotient s₁ → Quotient s₂ := Quot.map f h @[simp] @@ -702,15 +704,7 @@ theorem map'_mk'' (f : α → β) (h) (x : α) : rfl /-- A version of `Quotient.map₂` using curly braces and unification. -/ -protected def map₂' (f : α → β → γ) (h : (s₁.r ⇒ s₂.r ⇒ s₃.r) f f) : - Quotient s₁ → Quotient s₂ → Quotient s₃ := - Quotient.map₂ f h - -@[simp] -theorem map₂'_mk'' (f : α → β → γ) (h) (x : α) : - (Quotient.mk'' x : Quotient s₁).map₂' f h = - (Quotient.map' (f x) (h (Setoid.refl x)) : Quotient s₂ → Quotient s₃) := - rfl +@[deprecated (since := "2024-12-01")] protected alias map₂' := Quotient.map₂ theorem exact' {a b : α} : (Quotient.mk'' a : Quotient s₁) = Quotient.mk'' b → s₁ a b := diff --git a/Mathlib/Data/Setoid/Basic.lean b/Mathlib/Data/Setoid/Basic.lean index bbaa5b239fbfe..d860f84f98bd6 100644 --- a/Mathlib/Data/Setoid/Basic.lean +++ b/Mathlib/Data/Setoid/Basic.lean @@ -133,7 +133,7 @@ equivalence relations. -/ @[simps] def prodQuotientEquiv (r : Setoid α) (s : Setoid β) : Quotient r × Quotient s ≃ Quotient (r.prod s) where - toFun := fun (x, y) ↦ Quotient.map₂' Prod.mk (fun _ _ hx _ _ hy ↦ ⟨hx, hy⟩) x y + toFun := fun (x, y) ↦ Quotient.map₂ Prod.mk (fun _ _ hx _ _ hy ↦ ⟨hx, hy⟩) x y invFun := fun q ↦ Quotient.liftOn' q (fun xy ↦ (Quotient.mk'' xy.1, Quotient.mk'' xy.2)) fun x y hxy ↦ Prod.ext (by simpa using hxy.1) (by simpa using hxy.2) left_inv := fun q ↦ by diff --git a/Mathlib/GroupTheory/Congruence/Defs.lean b/Mathlib/GroupTheory/Congruence/Defs.lean index 0026330ad4b73..9ecaf70c760a9 100644 --- a/Mathlib/GroupTheory/Congruence/Defs.lean +++ b/Mathlib/GroupTheory/Congruence/Defs.lean @@ -281,7 +281,7 @@ protected theorem eq {a b : M} : (a : c.Quotient) = (b : c.Quotient) ↔ c a b : @[to_additive "The addition induced on the quotient by an additive congruence relation on a type with an addition."] instance hasMul : Mul c.Quotient := - ⟨Quotient.map₂' (· * ·) fun _ _ h1 _ _ h2 => c.mul h1 h2⟩ + ⟨Quotient.map₂ (· * ·) fun _ _ h1 _ _ h2 => c.mul h1 h2⟩ /-- The kernel of the quotient map induced by a congruence relation `c` equals `c`. -/ @[to_additive (attr := simp) "The kernel of the quotient map induced by an additive congruence @@ -739,7 +739,7 @@ instance hasInv : Inv c.Quotient := @[to_additive "The subtraction induced on the quotient by an additive congruence relation on a type with a subtraction."] instance hasDiv : Div c.Quotient := - ⟨(Quotient.map₂' (· / ·)) fun _ _ h₁ _ _ h₂ => c.div h₁ h₂⟩ + ⟨(Quotient.map₂ (· / ·)) fun _ _ h₁ _ _ h₂ => c.div h₁ h₂⟩ /-- The integer scaling induced on the quotient by a congruence relation on a type with a subtraction. -/ diff --git a/Mathlib/LinearAlgebra/Projectivization/Constructions.lean b/Mathlib/LinearAlgebra/Projectivization/Constructions.lean index fd7470292b540..cdbb83b605a70 100644 --- a/Mathlib/LinearAlgebra/Projectivization/Constructions.lean +++ b/Mathlib/LinearAlgebra/Projectivization/Constructions.lean @@ -67,7 +67,7 @@ variable [DecidableEq F] /-- Cross product on the projective plane. -/ def cross : ℙ F (Fin 3 → F) → ℙ F (Fin 3 → F) → ℙ F (Fin 3 → F) := - Quotient.map₂' (fun v w ↦ if h : crossProduct v.1 w.1 = 0 then v else ⟨crossProduct v.1 w.1, h⟩) + Quotient.map₂ (fun v w ↦ if h : crossProduct v.1 w.1 = 0 then v else ⟨crossProduct v.1 w.1, h⟩) (fun _ _ ⟨a, ha⟩ _ _ ⟨b, hb⟩ ↦ by simp_rw [← ha, ← hb, LinearMap.map_smul_of_tower, LinearMap.smul_apply, smul_smul, mul_comm b a, smul_eq_zero_iff_eq] diff --git a/Mathlib/Order/Filter/Germ/Basic.lean b/Mathlib/Order/Filter/Germ/Basic.lean index b7df9bc2a0c66..fe099e6b204fa 100644 --- a/Mathlib/Order/Filter/Germ/Basic.lean +++ b/Mathlib/Order/Filter/Germ/Basic.lean @@ -198,7 +198,7 @@ theorem map_map (op₁ : γ → δ) (op₂ : β → γ) (f : Germ l β) : /-- Lift a binary function `β → γ → δ` to a function `Germ l β → Germ l γ → Germ l δ`. -/ def map₂ (op : β → γ → δ) : Germ l β → Germ l γ → Germ l δ := - Quotient.map₂' (fun f g x => op (f x) (g x)) fun f f' Hf g g' Hg => + Quotient.map₂ (fun f g x => op (f x) (g x)) fun f f' Hf g g' Hg => Hg.mp <| Hf.mono fun x Hf Hg => by simp only [Hf, Hg] @[simp] diff --git a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean index 82a5c54c0a72c..1b14cfce13ed0 100644 --- a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean +++ b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean @@ -363,7 +363,7 @@ instance hasPow : Pow (HomogeneousLocalization 𝒜 x) ℕ where instance : Add (HomogeneousLocalization 𝒜 x) where add := - Quotient.map₂' (· + ·) + Quotient.map₂ (· + ·) fun c1 c2 (h : Localization.mk _ _ = Localization.mk _ _) c3 c4 (h' : Localization.mk _ _ = Localization.mk _ _) => by change Localization.mk _ _ = Localization.mk _ _ @@ -376,7 +376,7 @@ instance : Sub (HomogeneousLocalization 𝒜 x) where sub z1 z2 := z1 + -z2 instance : Mul (HomogeneousLocalization 𝒜 x) where mul := - Quotient.map₂' (· * ·) + Quotient.map₂ (· * ·) fun c1 c2 (h : Localization.mk _ _ = Localization.mk _ _) c3 c4 (h' : Localization.mk _ _ = Localization.mk _ _) => by change Localization.mk _ _ = Localization.mk _ _ diff --git a/Mathlib/Topology/Algebra/SeparationQuotient/Basic.lean b/Mathlib/Topology/Algebra/SeparationQuotient/Basic.lean index 1f2427e29b56f..6790b8c087127 100644 --- a/Mathlib/Topology/Algebra/SeparationQuotient/Basic.lean +++ b/Mathlib/Topology/Algebra/SeparationQuotient/Basic.lean @@ -85,7 +85,7 @@ variable {M : Type*} [TopologicalSpace M] @[to_additive] instance instMul [Mul M] [ContinuousMul M] : Mul (SeparationQuotient M) where - mul := Quotient.map₂' (· * ·) fun _ _ h₁ _ _ h₂ ↦ Inseparable.mul h₁ h₂ + mul := Quotient.map₂ (· * ·) fun _ _ h₁ _ _ h₂ ↦ Inseparable.mul h₁ h₂ @[to_additive (attr := simp)] theorem mk_mul [Mul M] [ContinuousMul M] (a b : M) : mk (a * b) = mk a * mk b := rfl @@ -168,7 +168,7 @@ instance instInvOneClass [InvOneClass G] [ContinuousInv G] : @[to_additive] instance instDiv [Div G] [ContinuousDiv G] : Div (SeparationQuotient G) where - div := Quotient.map₂' (· / ·) fun _ _ h₁ _ _ h₂ ↦ (Inseparable.prod h₁ h₂).map continuous_div' + div := Quotient.map₂ (· / ·) fun _ _ h₁ _ _ h₂ ↦ (Inseparable.prod h₁ h₂).map continuous_div' @[to_additive (attr := simp)] theorem mk_div [Div G] [ContinuousDiv G] (x y : G) : mk (x / y) = mk x / mk y := rfl From 2bd6fd41b873e6f019f86afd928e58d12a53f8e5 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Tue, 3 Dec 2024 15:54:17 +0000 Subject: [PATCH 497/829] chore: golf isRoot_of_isRoot_of_dvd_derivative_mul (#19586) --- Mathlib/Algebra/Polynomial/FieldDivision.lean | 9 ++++----- 1 file changed, 4 insertions(+), 5 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/FieldDivision.lean b/Mathlib/Algebra/Polynomial/FieldDivision.lean index bc0f49ad785ba..bb20dc5de2f78 100644 --- a/Mathlib/Algebra/Polynomial/FieldDivision.lean +++ b/Mathlib/Algebra/Polynomial/FieldDivision.lean @@ -175,11 +175,10 @@ See `isRoot_of_isRoot_iff_dvd_derivative_mul` -/ theorem isRoot_of_isRoot_of_dvd_derivative_mul [CharZero R] {f g : R[X]} (hf0 : f ≠ 0) (hfd : f ∣ f.derivative * g) {a : R} (haf : f.IsRoot a) : g.IsRoot a := by rcases hfd with ⟨r, hr⟩ - by_cases hdf0 : derivative f = 0 - · rw [eq_C_of_derivative_eq_zero hdf0] at haf hf0 - simp only [eval_C, derivative_C, zero_mul, dvd_zero, iff_true, IsRoot.def] at haf - rw [haf, map_zero] at hf0 - exact (hf0 rfl).elim + have hdf0 : derivative f ≠ 0 := by + contrapose! haf + rw [eq_C_of_derivative_eq_zero haf] at hf0 ⊢ + exact not_isRoot_C _ _ <| C_ne_zero.mp hf0 by_contra hg have hdfg0 : f.derivative * g ≠ 0 := mul_ne_zero hdf0 (by rintro rfl; simp at hg) have hr' := congr_arg (rootMultiplicity a) hr From a58d16068076a95415539d6b724b71dbba946f60 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Tue, 3 Dec 2024 16:13:01 +0000 Subject: [PATCH 498/829] chore: Generalize universes in `RingTheory/Filtration` (#19677) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib/RingTheory/Filtration.lean | 7 ++----- 1 file changed, 2 insertions(+), 5 deletions(-) diff --git a/Mathlib/RingTheory/Filtration.lean b/Mathlib/RingTheory/Filtration.lean index 49f369df496ee..ea599e0713283 100644 --- a/Mathlib/RingTheory/Filtration.lean +++ b/Mathlib/RingTheory/Filtration.lean @@ -40,10 +40,7 @@ This file contains the definitions and basic results around (stable) `I`-filtrat -/ - -universe u v - -variable {R M : Type u} [CommRing R] [AddCommGroup M] [Module R M] (I : Ideal R) +variable {R M : Type*} [CommRing R] [AddCommGroup M] [Module R M] (I : Ideal R) open Polynomial @@ -52,7 +49,7 @@ open scoped Polynomial /-- An `I`-filtration on the module `M` is a sequence of decreasing submodules `N i` such that `I • (N i) ≤ N (i + 1)`. Note that we do not require the filtration to start from `⊤`. -/ @[ext] -structure Ideal.Filtration (M : Type u) [AddCommGroup M] [Module R M] where +structure Ideal.Filtration (M : Type*) [AddCommGroup M] [Module R M] where N : ℕ → Submodule R M mono : ∀ i, N (i + 1) ≤ N i smul_le : ∀ i, I • N i ≤ N (i + 1) From 14e2a6b76e28622f84380169ed4f37b13d596a1c Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Tue, 3 Dec 2024 16:48:42 +0000 Subject: [PATCH 499/829] feat(MvPolynomial/PDeriv): Euler's (weighted) homogeneous identity (#13946) If $f$ is a weighted homogeneous polynomial of degree $n$ with a finite set of variables $\\{X_i \mid i \in \sigma\\}$, then $n\cdot f = \sum_{i \in \sigma} w_i X_i \frac{\partial f}{\partial X_i}$, where $w_i$ are the weights. Co-authored-by: Junyan Xu --- Mathlib.lean | 1 + Mathlib/Algebra/MvPolynomial/Degrees.lean | 9 +++ Mathlib/Algebra/MvPolynomial/PDeriv.lean | 7 ++ Mathlib/Data/Finsupp/Order.lean | 4 + Mathlib/Data/Finsupp/Weight.lean | 7 ++ .../MvPolynomial/EulerIdentity.lean | 73 +++++++++++++++++++ .../MvPolynomial/WeightedHomogeneous.lean | 6 ++ 7 files changed, 107 insertions(+) create mode 100644 Mathlib/RingTheory/MvPolynomial/EulerIdentity.lean diff --git a/Mathlib.lean b/Mathlib.lean index d9fc3061091cf..b24ff15bb5cab 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4431,6 +4431,7 @@ import Mathlib.RingTheory.MaximalSpectrum import Mathlib.RingTheory.Multiplicity import Mathlib.RingTheory.MvPolynomial import Mathlib.RingTheory.MvPolynomial.Basic +import Mathlib.RingTheory.MvPolynomial.EulerIdentity import Mathlib.RingTheory.MvPolynomial.FreeCommRing import Mathlib.RingTheory.MvPolynomial.Homogeneous import Mathlib.RingTheory.MvPolynomial.Ideal diff --git a/Mathlib/Algebra/MvPolynomial/Degrees.lean b/Mathlib/Algebra/MvPolynomial/Degrees.lean index 45521d9bd0bb7..5227f6ddbbe60 100644 --- a/Mathlib/Algebra/MvPolynomial/Degrees.lean +++ b/Mathlib/Algebra/MvPolynomial/Degrees.lean @@ -500,6 +500,15 @@ theorem coeff_eq_zero_of_totalDegree_lt {f : MvPolynomial σ R} {d : σ →₀ exact lt_irrefl _ · exact lt_of_le_of_lt (Nat.zero_le _) h +theorem totalDegree_eq_zero_iff_eq_C {p : MvPolynomial σ R} : + p.totalDegree = 0 ↔ p = C (p.coeff 0) := by + constructor <;> intro h + · ext m; classical rw [coeff_C]; split_ifs with hm; · rw [← hm] + apply coeff_eq_zero_of_totalDegree_lt; rw [h] + exact Finset.sum_pos (fun i hi ↦ Nat.pos_of_ne_zero <| Finsupp.mem_support_iff.mp hi) + (Finsupp.support_nonempty_iff.mpr <| Ne.symm hm) + · rw [h, totalDegree_C] + theorem totalDegree_rename_le (f : σ → τ) (p : MvPolynomial σ R) : (rename f p).totalDegree ≤ p.totalDegree := Finset.sup_le fun b => by diff --git a/Mathlib/Algebra/MvPolynomial/PDeriv.lean b/Mathlib/Algebra/MvPolynomial/PDeriv.lean index 9c6e2c3c43e7c..5cafce58bab3e 100644 --- a/Mathlib/Algebra/MvPolynomial/PDeriv.lean +++ b/Mathlib/Algebra/MvPolynomial/PDeriv.lean @@ -72,6 +72,13 @@ theorem pderiv_monomial {i : σ} : · rw [Finsupp.not_mem_support_iff] at hi; simp [hi] · simp +lemma X_mul_pderiv_monomial {i : σ} {m : σ →₀ ℕ} {r : R} : + X i * pderiv i (monomial m r) = m i • monomial m r := by + rw [pderiv_monomial, X, monomial_mul, smul_monomial] + by_cases h : m i = 0 + · simp_rw [h, Nat.cast_zero, mul_zero, zero_smul, monomial_zero] + rw [one_mul, mul_comm, nsmul_eq_mul, add_comm, sub_add_single_one_cancel h] + theorem pderiv_C {i : σ} : pderiv i (C a) = 0 := derivation_C _ _ diff --git a/Mathlib/Data/Finsupp/Order.lean b/Mathlib/Data/Finsupp/Order.lean index b06f454769ac2..ad9bfbd4b5485 100644 --- a/Mathlib/Data/Finsupp/Order.lean +++ b/Mathlib/Data/Finsupp/Order.lean @@ -322,6 +322,10 @@ theorem add_sub_single_one {a : ι} {u u' : ι →₀ ℕ} (h : u' a ≠ 0) : u + (u' - single a 1) = u + u' - single a 1 := (add_tsub_assoc_of_le (single_le_iff.mpr <| Nat.one_le_iff_ne_zero.mpr h) _).symm +lemma sub_add_single_one_cancel {u : ι →₀ ℕ} {i : ι} (h : u i ≠ 0) : + u - single i 1 + single i 1 = u := by + rw [sub_single_one_add h, add_tsub_cancel_right] + end Nat end Finsupp diff --git a/Mathlib/Data/Finsupp/Weight.lean b/Mathlib/Data/Finsupp/Weight.lean index 2278beedc1aec..7b07c03ebf167 100644 --- a/Mathlib/Data/Finsupp/Weight.lean +++ b/Mathlib/Data/Finsupp/Weight.lean @@ -105,6 +105,13 @@ theorem NonTorsionWeight.ne_zero [NonTorsionWeight w] (s : σ) : apply Nat.zero_ne_one.symm exact NonTorsionWeight.eq_zero_of_smul_eq_zero h +variable {w} in +lemma weight_sub_single_add {f : σ →₀ ℕ} {i : σ} (hi : f i ≠ 0) : + (f - single i 1).weight w + w i = f.weight w := by + conv_rhs => rw [← sub_add_single_one_cancel hi, weight_apply] + rw [sum_add_index', sum_single_index, one_smul, weight_apply] + exacts [zero_smul .., fun _ ↦ zero_smul .., fun _ _ _ ↦ add_smul ..] + end AddCommMonoid section OrderedAddCommMonoid diff --git a/Mathlib/RingTheory/MvPolynomial/EulerIdentity.lean b/Mathlib/RingTheory/MvPolynomial/EulerIdentity.lean new file mode 100644 index 0000000000000..80d599cc70bfb --- /dev/null +++ b/Mathlib/RingTheory/MvPolynomial/EulerIdentity.lean @@ -0,0 +1,73 @@ +/- +Copyright (c) 2024 Junyan Xu. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Junyan Xu +-/ +import Mathlib.Algebra.MvPolynomial.PDeriv +import Mathlib.RingTheory.MvPolynomial.Homogeneous + +/-! +# Euler's homogeneous identity + +## Main resutls + +* `IsHomogeneous.sum_X_mul_pderiv`: Euler's identity for homogeneous polynomials: + for a multivariate homogeneous polynomial, + the product of each variable with the derivative with respect to that variable + sums up to the degree times the polynomial. +* `IsWeightedHomogeneous.sum_weight_X_mul_pderiv`: the weighted version of Euler's identity. +-/ + +namespace MvPolynomial + +open Finsupp + +variable {R σ M : Type*} [CommSemiring R] {φ : MvPolynomial σ R} + +protected lemma IsWeightedHomogeneous.pderiv [AddCancelCommMonoid M] {w : σ → M} {n n' : M} {i : σ} + (h : φ.IsWeightedHomogeneous w n) (h' : n' + w i = n) : + (pderiv i φ).IsWeightedHomogeneous w n' := by + rw [← mem_weightedHomogeneousSubmodule, weightedHomogeneousSubmodule_eq_finsupp_supported, + Finsupp.supported_eq_span_single] at h + refine Submodule.span_induction ?_ ?_ (fun p q _ _ hp hq ↦ ?_) (fun r p _ h ↦ ?_) h + · rintro _ ⟨m, hm, rfl⟩ + simp_rw [single_eq_monomial, pderiv_monomial, one_mul] + by_cases hi : m i = 0 + · rw [hi, Nat.cast_zero, monomial_zero]; apply isWeightedHomogeneous_zero + convert isWeightedHomogeneous_monomial .. + rw [← add_right_cancel_iff (a := w i), h', ← hm, weight_sub_single_add hi] + · rw [map_zero]; apply isWeightedHomogeneous_zero + · rw [map_add]; exact hp.add hq + · rw [(pderiv i).map_smul]; exact (weightedHomogeneousSubmodule ..).smul_mem _ h + +protected lemma IsHomogeneous.pderiv {n : ℕ} {i : σ} (h : φ.IsHomogeneous n) : + (pderiv i φ).IsHomogeneous (n - 1) := by + obtain _ | n := n + · rw [← totalDegree_zero_iff_isHomogeneous, totalDegree_eq_zero_iff_eq_C] at h + rw [h, pderiv_C]; apply isHomogeneous_zero + · exact IsWeightedHomogeneous.pderiv h rfl + +variable [Fintype σ] {n : ℕ} + +open Finset in +/-- Euler's identity for weighted homogeneous polynomials. -/ +theorem IsWeightedHomogeneous.sum_weight_X_mul_pderiv {w : σ → ℕ} + (h : φ.IsWeightedHomogeneous w n) : ∑ i : σ, w i • (X i * pderiv i φ) = n • φ := by + rw [← mem_weightedHomogeneousSubmodule, weightedHomogeneousSubmodule_eq_finsupp_supported, + supported_eq_span_single] at h + refine Submodule.span_induction ?_ ?_ (fun p q _ _ hp hq ↦ ?_) (fun r p _ h ↦ ?_) h + · rintro _ ⟨m, hm, rfl⟩ + simp_rw [single_eq_monomial, X_mul_pderiv_monomial, smul_smul, ← sum_smul, mul_comm (w _)] + congr + rwa [Set.mem_setOf, weight_apply, sum_fintype] at hm + intro; apply zero_smul + · simp + · simp_rw [map_add, left_distrib, smul_add, sum_add_distrib, hp, hq] + · simp_rw [(pderiv _).map_smul, nsmul_eq_mul, mul_smul_comm, ← Finset.smul_sum, ← nsmul_eq_mul, h] + +/-- Euler's identity for homogeneous polynomials. -/ +theorem IsHomogeneous.sum_X_mul_pderiv (h : φ.IsHomogeneous n) : + ∑ i : σ, X i * pderiv i φ = n • φ := by + simp_rw [← h.sum_weight_X_mul_pderiv, Pi.one_apply, one_smul] + +end MvPolynomial diff --git a/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean b/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean index 18bd357e9e117..c0b99872fcf13 100644 --- a/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean +++ b/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean @@ -251,6 +251,12 @@ theorem mul {w : σ → M} (hφ : IsWeightedHomogeneous w φ m) (hψ : IsWeighte IsWeightedHomogeneous w (φ * ψ) (m + n) := weightedHomogeneousSubmodule_mul w m n <| Submodule.mul_mem_mul hφ hψ +theorem pow {w : σ → M} (hφ : IsWeightedHomogeneous w φ m) (n : ℕ) : + IsWeightedHomogeneous w (φ ^ n) (n • m) := by + induction n with + | zero => rw [pow_zero, zero_smul]; exact isWeightedHomogeneous_one R w + | succ n ih => rw [pow_succ, succ_nsmul]; exact ih.mul hφ + /-- A product of weighted homogeneous polynomials is weighted homogeneous, with weighted degree equal to the sum of the weighted degrees. -/ theorem prod {ι : Type*} (s : Finset ι) (φ : ι → MvPolynomial σ R) (n : ι → M) {w : σ → M} : From da9df8db61024281c6beb7853df1ccec445ba61b Mon Sep 17 00:00:00 2001 From: Alex Meiburg Date: Tue, 3 Dec 2024 19:01:43 +0000 Subject: [PATCH 500/829] feat: `List.destutter` on cotransitive relations. (#9082) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit `List.destutter` behaves nicely when the condition is a _cotransitive_ relation, and even better when it's a _coequivalence_ relation. This proves that behavior. Co-authored-by: Alex Meiburg Co-authored-by: Yaël Dillies --- Mathlib/Data/List/Destutter.lean | 117 ++++++++++++++++++++++++++++++- Mathlib/Data/List/GetD.lean | 10 +++ Mathlib/Order/Basic.lean | 4 ++ 3 files changed, 130 insertions(+), 1 deletion(-) diff --git a/Mathlib/Data/List/Destutter.lean b/Mathlib/Data/List/Destutter.lean index df3a1ad3b5778..f021d2425e7a0 100644 --- a/Mathlib/Data/List/Destutter.lean +++ b/Mathlib/Data/List/Destutter.lean @@ -25,8 +25,11 @@ Note that we make no guarantees of being the longest sublist with this property; adjacent, chain, duplicates, remove, list, stutter, destutter -/ +open Function -variable {α : Type*} (l : List α) (R : α → α → Prop) [DecidableRel R] {a b : α} +variable {α β : Type*} (l l₁ l₂ : List α) (R : α → α → Prop) [DecidableRel R] {a b : α} + +variable {R₂ : β → β → Prop} [DecidableRel R₂] namespace List @@ -148,4 +151,116 @@ theorem destutter_eq_nil : ∀ {l : List α}, destutter R l = [] ↔ l = [] | [] => Iff.rfl | _ :: l => ⟨fun h => absurd h <| l.destutter'_ne_nil R, fun h => nomatch h⟩ +variable {R} + +/-- For a relation-preserving map, `destutter` commutes with `map`. -/ +theorem map_destutter {f : α → β} : ∀ {l : List α}, (∀ a ∈ l, ∀ b ∈ l, R a b ↔ R₂ (f a) (f b)) → + (l.destutter R).map f = (l.map f).destutter R₂ + | [], hl => by simp + | [a], hl => by simp + | a :: b :: l, hl => by + have := hl a (by simp) b (by simp) + simp_rw [map_cons, destutter_cons_cons, ← this] + by_cases hr : R a b <;> + simp [hr, ← destutter_cons', map_destutter fun c hc d hd ↦ hl _ (cons_subset_cons _ + (subset_cons_self _ _) hc) _ (cons_subset_cons _ (subset_cons_self _ _) hd), + map_destutter fun c hc d hd ↦ hl _ (subset_cons_self _ _ hc) _ (subset_cons_self _ _ hd)] + +/-- For a injective function `f`, `destutter' (·≠·)` commutes with `map f`. -/ +theorem map_destutter_ne {f : α → β} (h : Injective f) [DecidableEq α] [DecidableEq β] : + (l.destutter (·≠·)).map f = (l.map f).destutter (·≠·) := + map_destutter fun _ _ _ _ ↦ h.ne_iff.symm + +/-- `destutter'` on a relation like ≠ or <, whose negation is transitive, has length monotone +under a `¬R` changing of the first element. -/ +theorem length_destutter'_cotrans_ge [i : IsTrans α Rᶜ] : + ∀ {a} {l : List α}, ¬R b a → (l.destutter' R b).length ≤ (l.destutter' R a).length + | a, [], hba => by simp + | a, c :: l, hba => by + by_cases hbc : R b c + case pos => + have hac : ¬Rᶜ a c := (mt (_root_.trans hba)) (not_not.2 hbc) + simp_rw [destutter', if_pos (not_not.1 hac), if_pos hbc, length_cons, le_refl] + case neg => + simp only [destutter', if_neg hbc] + by_cases hac : R a c + case pos => + simp only [if_pos hac, length_cons] + exact Nat.le_succ_of_le (length_destutter'_cotrans_ge hbc) + case neg => + simp only [if_neg hac] + exact length_destutter'_cotrans_ge hba + +/-- `List.destutter'` on a relation like `≠`, whose negation is an equivalence, gives the same +length if the first elements are not related. -/ +theorem length_destutter'_congr [IsEquiv α Rᶜ] (hab : ¬R a b) : + (l.destutter' R a).length = (l.destutter' R b).length := + (length_destutter'_cotrans_ge hab).antisymm <| length_destutter'_cotrans_ge (symm hab : Rᶜ b a) + +/-- `List.destutter'` on a relation like ≠, whose negation is an equivalence, has length + monotonic under List.cons -/ +/- +TODO: Replace this lemma by the more general version: +theorem Sublist.length_destutter'_mono [IsEquiv α Rᶜ] (h : a :: l₁ <+ b :: l₂) : + (List.destutter' R a l₁).length ≤ (List.destutter' R b l₂).length +-/ +theorem le_length_destutter'_cons [IsEquiv α Rᶜ] : + ∀ {l : List α}, (l.destutter' R b).length ≤ ((b :: l).destutter' R a).length + | [] => by by_cases hab : (R a b) <;> simp_all [Nat.le_succ] + | c :: cs => by + by_cases hab : R a b + case pos => simp [destutter', if_pos hab, Nat.le_succ] + obtain hac | hac : R a c ∨ Rᶜ a c := em _ + · have hbc : ¬Rᶜ b c := mt (_root_.trans hab) (not_not.2 hac) + simp [destutter', if_pos hac, if_pos (not_not.1 hbc), if_neg hab] + · have hbc : ¬R b c := trans (symm hab) hac + simp only [destutter', if_neg hbc, if_neg hac, if_neg hab] + exact (length_destutter'_congr cs hab).ge + +/-- `List.destutter` on a relation like ≠, whose negation is an equivalence, has length +monotone under List.cons -/ +theorem length_destutter_le_length_destutter_cons [IsEquiv α Rᶜ] : + ∀ {l : List α}, (l.destutter R).length ≤ ((a :: l).destutter R).length + | [] => by simp [destutter] + | b :: l => le_length_destutter'_cons + +variable {l l₁ l₂} + +/-- `destutter ≠` has length monotone under `List.cons`. -/ +theorem length_destutter_ne_le_length_destutter_cons [DecidableEq α] : + (l.destutter (· ≠ ·)).length ≤ ((a :: l).destutter (· ≠ ·)).length := + length_destutter_le_length_destutter_cons + +/-- `destutter` of relations like `≠`, whose negation is an equivalence relation, +gives a list of maximal length over any chain. + +In other words, `l.destutter R` is an `R`-chain sublist of `l`, and is at least as long as any other +`R`-chain sublist. -/ +lemma Chain'.length_le_length_destutter [IsEquiv α Rᶜ] : + ∀ {l₁ l₂ : List α}, l₁ <+ l₂ → l₁.Chain' R → l₁.length ≤ (l₂.destutter R).length + -- `l₁ := []`, `l₂ := []` + | [], [], _, _ => by simp + -- `l₁ := l₁`, `l₂ := a :: l₂` + | l₁, _, .cons (l₂ := l₂) a hl, hl₁ => + (hl₁.length_le_length_destutter hl).trans length_destutter_le_length_destutter_cons + -- `l₁ := [a]`, `l₂ := a :: l₂` + | _, _, .cons₂ (l₁ := []) (l₂ := l₁) a hl, hl₁ => by simp [Nat.one_le_iff_ne_zero] + -- `l₁ := a :: l₁`, `l₂ := a :: b :: l₂` + | _, _, .cons₂ a <| .cons (l₁ := l₁) (l₂ := l₂) b hl, hl₁ => by + by_cases hab : R a b + · simpa [destutter_cons_cons, hab] using hl₁.tail.length_le_length_destutter (hl.cons _) + · simpa [destutter_cons_cons, hab] using hl₁.length_le_length_destutter (hl.cons₂ _) + -- `l₁ := a :: b :: l₁`, `l₂ := a :: b :: l₂` + | _, _, .cons₂ a <| .cons₂ (l₁ := l₁) (l₂ := l₂) b hl, hl₁ => by + simpa [destutter_cons_cons, rel_of_chain_cons hl₁] + using hl₁.tail.length_le_length_destutter (hl.cons₂ _) + +/-- `destutter` of `≠` gives a list of maximal length over any chain. + +In other words, `l.destutter (· ≠ ·)` is a `≠`-chain sublist of `l`, and is at least as long as any +other `≠`-chain sublist. -/ +lemma Chain'.length_le_length_destutter_ne [DecidableEq α] (hl : l₁ <+ l₂) + (hl₁ : l₁.Chain' (· ≠ ·)) : l₁.length ≤ (l₂.destutter (· ≠ ·)).length := + hl₁.length_le_length_destutter hl + end List diff --git a/Mathlib/Data/List/GetD.lean b/Mathlib/Data/List/GetD.lean index 5aaf7611ee1bd..9e93d425fe171 100644 --- a/Mathlib/Data/List/GetD.lean +++ b/Mathlib/Data/List/GetD.lean @@ -49,6 +49,11 @@ theorem getD_eq_default {n : ℕ} (hn : l.length ≤ n) : l.getD n d = d := by · simp at hn · exact ih (Nat.le_of_succ_le_succ hn) +theorem getD_reverse {l : List α} (i) (h : i < length l) : + getD l.reverse i = getD l (l.length - 1 - i) := by + funext a + rwa [List.getD_eq_getElem?_getD, List.getElem?_reverse, ← List.getD_eq_getElem?_getD] + /-- An empty list can always be decidably checked for the presence of an element. Not an instance because it would clash with `DecidableEq α`. -/ def decidableGetDNilNe (a : α) : DecidablePred fun i : ℕ => getD ([] : List α) i a ≠ a := @@ -69,6 +74,11 @@ theorem getElem?_getD_replicate_default_eq (r n : ℕ) : (replicate r d)[n]?.get @[deprecated (since := "2024-06-12")] alias getD_replicate_default_eq := getElem?_getD_replicate_default_eq +theorem getD_replicate {y i n} (h : i < n) : + getD (replicate n x) i y = x := by + rw [getD_eq_getElem, getElem_replicate] + rwa [length_replicate] + 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 _ _)), diff --git a/Mathlib/Order/Basic.lean b/Mathlib/Order/Basic.lean index caa54da5b5024..8820922f5b0b4 100644 --- a/Mathlib/Order/Basic.lean +++ b/Mathlib/Order/Basic.lean @@ -762,6 +762,10 @@ theorem compl_le [LinearOrder α] : (· ≤ · : α → α → _)ᶜ = (· > ·) theorem compl_gt [LinearOrder α] : (· > · : α → α → _)ᶜ = (· ≤ ·) := by ext; simp [compl] theorem compl_ge [LinearOrder α] : (· ≥ · : α → α → _)ᶜ = (· < ·) := by ext; simp [compl] +instance Ne.instIsEquiv_compl : IsEquiv α (· ≠ ·)ᶜ := by + convert eq_isEquiv α + simp [compl] + /-! ### Order instances on the function space -/ From 4969d8c2185478d362c5b9eb8fd8684a3c80ee7c Mon Sep 17 00:00:00 2001 From: Alex Meiburg Date: Wed, 4 Dec 2024 02:21:57 +0000 Subject: [PATCH 501/829] feat: `List.destutter` on cotransitive relations. (#9082) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit `List.destutter` behaves nicely when the condition is a _cotransitive_ relation, and even better when it's a _coequivalence_ relation. This proves that behavior. Co-authored-by: Alex Meiburg Co-authored-by: Yaël Dillies From d3922baf4ea018963d5796482512b77bca4497a2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?=C3=93scar=20=C3=81lvarez?= Date: Wed, 4 Dec 2024 05:17:51 +0000 Subject: [PATCH 502/829] feat(GroupTheory/Coxeter): add auxiliary lemmas for alternating words and left inverse sequences (#19663) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add some auxiliary lemmas of alternating words and left inverse sequences of Coxeter groups that will be useful to prove the Strong Exchange Theorem (among other uses). They are mainly taken from pages 12-13 of [Combinatorics of Coxeter groups by Anders Bjorner , Francesco Brenti](https://link.springer.com/book/10.1007/3-540-27596-7). Co-authored-by: Óscar Co-authored-by: Óscar Álvarez --- Mathlib/GroupTheory/Coxeter/Basic.lean | 77 +++++++++++++++++++++- Mathlib/GroupTheory/Coxeter/Inversion.lean | 49 +++++++++++++- 2 files changed, 124 insertions(+), 2 deletions(-) diff --git a/Mathlib/GroupTheory/Coxeter/Basic.lean b/Mathlib/GroupTheory/Coxeter/Basic.lean index 1ab0f3e0bef58..c33678eadac13 100644 --- a/Mathlib/GroupTheory/Coxeter/Basic.lean +++ b/Mathlib/GroupTheory/Coxeter/Basic.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2024 Newell Jensen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Newell Jensen, Mitchell Lee +Authors: Newell Jensen, Mitchell Lee, Óscar Álvarez -/ import Mathlib.Algebra.Group.Subgroup.Pointwise import Mathlib.Algebra.Ring.Int.Parity @@ -411,6 +411,81 @@ theorem length_alternatingWord (i i' : B) (m : ℕ) : · dsimp [alternatingWord] · simpa [alternatingWord] using ih i' i +lemma getElem_alternatingWord (i j : B) (p k : ℕ) (hk : k < p) : + (alternatingWord i j p)[k]'(by simp; exact hk) = (if Even (p + k) then i else j) := by + revert k + induction p with + | zero => + intro k hk + simp only [not_lt_zero'] at hk + | succ n h => + intro k hk + simp_rw [alternatingWord_succ' i j n] + match k with + | 0 => + by_cases h2 : Even n + · simp only [h2, ↓reduceIte, getElem_cons_zero, add_zero, + (by simp [Even.add_one, h2] : ¬Even (n + 1))] + · simp only [h2, ↓reduceIte, getElem_cons_zero, add_zero, + Odd.add_one (Nat.not_even_iff_odd.mp h2)] + | k + 1 => + simp only [add_lt_add_iff_right] at hk h + simp only [getElem_cons_succ, h k hk] + ring_nf + have even_add_two (m : ℕ) : Even (2 + m) ↔ Even m := by + simp only [add_tsub_cancel_right, even_two, (Nat.even_sub (by omega : m ≤ 2 + m)).mp] + by_cases h_even : Even (n + k) + · rw [if_pos h_even] + rw [← even_add_two (n+k), ← Nat.add_assoc 2 n k] at h_even + rw [if_pos h_even] + · rw [if_neg h_even] + rw [← even_add_two (n+k), ← Nat.add_assoc 2 n k] at h_even + rw [if_neg h_even] + +lemma getElem_alternatingWord_swapIndices (i j : B) (p k : ℕ) (h : k + 1 < p) : + (alternatingWord i j p)[k+1]'(by simp; exact h) = + (alternatingWord j i p)[k]'(by simp [h]; omega) := by + rw [getElem_alternatingWord i j p (k+1) (by omega), getElem_alternatingWord j i p k (by omega)] + by_cases h_even : Even (p + k) + · rw [if_pos h_even, ← add_assoc] + simp only [ite_eq_right_iff, isEmpty_Prop, Nat.not_even_iff_odd, Even.add_one h_even, + IsEmpty.forall_iff] + · rw [if_neg h_even, ← add_assoc] + simp [Odd.add_one (Nat.not_even_iff_odd.mp h_even)] + +lemma listTake_alternatingWord (i j : B) (p k : ℕ) (h : k < 2 * p) : + List.take k (alternatingWord i j (2 * p)) = + if Even k then alternatingWord i j k else alternatingWord j i k := by + induction k with + | zero => + simp only [take_zero, even_zero, ↓reduceIte, alternatingWord] + | succ k h' => + have hk : k < 2 * p := by omega + apply h' at hk + by_cases h_even : Even k + · simp only [h_even, ↓reduceIte] at hk + simp only [Nat.not_even_iff_odd.mpr (Even.add_one h_even), ↓reduceIte] + rw [← List.take_concat_get _ _ (by simp[h]; omega), alternatingWord_succ, ← hk] + apply congr_arg + rw [getElem_alternatingWord i j (2*p) k (by omega)] + simp [(by apply Nat.even_add.mpr; simp[h_even]: Even (2 * p + k))] + · simp only [h_even, ↓reduceIte] at hk + simp only [(by simp at h_even; exact Odd.add_one h_even : Even (k + 1)), ↓reduceIte] + rw [← List.take_concat_get _ _ (by simp[h]; omega), alternatingWord_succ, hk] + apply congr_arg + rw [getElem_alternatingWord i j (2*p) k (by omega)] + simp [(by apply Nat.odd_add.mpr; simp[h_even]: Odd (2 * p + k))] + +lemma listTake_succ_alternatingWord (i j : B) (p : ℕ) (k : ℕ) (h : k + 1 < 2 * p) : + List.take (k + 1) (alternatingWord i j (2 * p)) = + i :: (List.take k (alternatingWord j i (2 * p))) := by + rw [listTake_alternatingWord j i p k (by omega), listTake_alternatingWord i j p (k+1) h] + + by_cases h_even : Even k + · simp [h_even, Nat.not_even_iff_odd.mpr (Even.add_one h_even), alternatingWord_succ', h_even] + · simp [h_even, (by simp at h_even; exact Odd.add_one h_even: Even (k + 1)), + alternatingWord_succ', h_even] + theorem prod_alternatingWord_eq_mul_pow (i i' : B) (m : ℕ) : π (alternatingWord i i' m) = (if Even m then 1 else s i') * (s i * s i') ^ (m / 2) := by induction' m with m ih diff --git a/Mathlib/GroupTheory/Coxeter/Inversion.lean b/Mathlib/GroupTheory/Coxeter/Inversion.lean index 6afee6ebfd411..b757715af2158 100644 --- a/Mathlib/GroupTheory/Coxeter/Inversion.lean +++ b/Mathlib/GroupTheory/Coxeter/Inversion.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2024 Mitchell Lee. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mitchell Lee +Authors: Mitchell Lee, Óscar Álvarez -/ import Mathlib.GroupTheory.Coxeter.Length import Mathlib.Data.List.GetD @@ -255,6 +255,13 @@ theorem getD_rightInvSeq (ω : List B) (j : ℕ) : · simp only [getD_eq_getElem?_getD, get?_eq_getElem?] at ih simp [getD_cons_succ, ih j'] +lemma getElem_rightInvSeq (ω : List B) (j : ℕ) (h : j < ω.length) : + (ris ω)[j]'(by simp[h]) = + (π (ω.drop (j + 1)))⁻¹ + * (Option.map (cs.simple) (ω.get? j)).getD 1 + * π (ω.drop (j + 1)) := by + rw [← List.getD_eq_getElem (ris ω) 1, getD_rightInvSeq] + theorem getD_leftInvSeq (ω : List B) (j : ℕ) : (lis ω).getD j 1 = π (ω.take j) @@ -271,6 +278,12 @@ theorem getD_leftInvSeq (ω : List B) (j : ℕ) : rw [ih j'] simp [← mul_assoc, wordProd_cons] +lemma getElem_leftInvSeq (ω : List B) (j : ℕ) (h : j < ω.length) : + (lis ω)[j]'(by simp[h]) = + cs.wordProd (List.take j ω) * s ω[j] * (cs.wordProd (List.take j ω))⁻¹ := by + rw [← List.getD_eq_getElem (lis ω) 1, getD_leftInvSeq] + simp [h] + theorem getD_rightInvSeq_mul_self (ω : List B) (j : ℕ) : ((ris ω).getD j 1) * ((ris ω).getD j 1) = 1 := by simp_rw [getD_rightInvSeq, mul_assoc] @@ -439,4 +452,38 @@ theorem IsReduced.nodup_leftInvSeq {ω : List B} (rω : cs.IsReduced ω) : List. apply nodup_rightInvSeq rwa [isReduced_reverse] +lemma getElem_succ_leftInvSeq_alternatingWord + (i j : B) (p k : ℕ) (h : k + 1 < 2 * p) : + (lis (alternatingWord i j (2 * p)))[k + 1]'(by simp; exact h) = + MulAut.conj (s i) ((lis (alternatingWord j i (2 * p)))[k]'(by simp; linarith)) := by + rw [cs.getElem_leftInvSeq (alternatingWord i j (2 * p)) (k + 1) (by simp[h]), + cs.getElem_leftInvSeq (alternatingWord j i (2 * p)) k (by simp[h]; omega)] + simp only [MulAut.conj, listTake_succ_alternatingWord i j p k h, cs.wordProd_cons, mul_assoc, + mul_inv_rev, inv_simple, MonoidHom.coe_mk, OneHom.coe_mk, MulEquiv.coe_mk, Equiv.coe_fn_mk, + mul_right_inj, mul_left_inj] + rw [getElem_alternatingWord_swapIndices i j (2 * p) k] + omega + +theorem getElem_leftInvSeq_alternatingWord + (i j : B) (p k : ℕ) (h : k < 2 * p) : + (lis (alternatingWord i j (2 * p)))[k]'(by simp; linarith) = + π alternatingWord j i (2 * k + 1) := by + revert i j + induction k with + | zero => + intro i j + simp only [CoxeterSystem.getElem_leftInvSeq cs (alternatingWord i j (2 * p)) 0 (by simp [h]), + take_zero, wordProd_nil, one_mul, inv_one, mul_one, alternatingWord, concat_eq_append, + nil_append, wordProd_singleton] + apply congr_arg + simp only [getElem_alternatingWord i j (2 * p) 0 (by simp [h]), add_zero, even_two, + Even.mul_right, ↓reduceIte] + | succ k hk => + intro i j + simp only [getElem_succ_leftInvSeq_alternatingWord cs i j p k h, hk (by omega), + MulAut.conj_apply, inv_simple, alternatingWord_succ' j i, even_two, Even.mul_right, + ↓reduceIte, wordProd_cons] + rw [(by ring: 2 * (k + 1) = 2 * k + 1 + 1), alternatingWord_succ j i, wordProd_concat] + simp [mul_assoc] + end CoxeterSystem From 61bb98a3995e0fe19ee9cc7e2be11fb5533aa26e Mon Sep 17 00:00:00 2001 From: Mitchell Lee Date: Wed, 4 Dec 2024 05:49:32 +0000 Subject: [PATCH 503/829] feat: the Chebyshev polynomials C and S (#13195) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We define the rescaled Chebyshev polynomials $C_n$ and $S_n$ (also known as the Vieta–Lucas and Vieta–Fibonacci polynomials, respectively). They are related to the Chebyshev polynomials $T_n$ and $U_n$ by the formulas $C_n(2x) = 2T_n(x)$ and $S_n(2x) = U_n(x)$. Most theorems about $T_n$ and $U_n$ have analogues involving $C_n$ and $S_n$. We prove that $C_n$ and $S_n$ are special cases of the Dickson polynomials (though unlike general Dickson polynomials, they are defined for every integer $n$, not just natural numbers). These polynomials are necessary to state a formula for $(r_1 r_2)^m v$, where $v \in V$ is an element of a module, $r_1, r_2 \in GL(V)$ are reflections, and $m$ is an integer. The formula will be used to define and construct reflection representations of a Coxeter group over an arbitrary commutative ring, not necessarily having an invertible 2. See #13291. Co-authored-by: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> --- .../Trigonometric/Chebyshev.lean | 24 ++ Mathlib/RingTheory/Polynomial/Chebyshev.lean | 259 +++++++++++++++++- Mathlib/RingTheory/Polynomial/Dickson.lean | 44 ++- 3 files changed, 303 insertions(+), 24 deletions(-) diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Chebyshev.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Chebyshev.lean index 9e2609f0ac9b5..23c49de796496 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Chebyshev.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Chebyshev.lean @@ -94,6 +94,18 @@ theorem U_complex_cos (n : ℤ) : (U ℂ n).eval (cos θ) * sin θ = sin ((n + 1 push_cast ring_nf +/-- The `n`-th rescaled Chebyshev polynomial of the first kind (Vieta–Lucas polynomial) evaluates on +`2 * cos θ` to the value `2 * cos (n * θ)`. -/ +@[simp] +theorem C_two_mul_complex_cos (n : ℤ) : (C ℂ n).eval (2 * cos θ) = 2 * cos (n * θ) := by + simp [C_eq_two_mul_T_comp_half_mul_X] + +/-- The `n`-th rescaled Chebyshev polynomial of the second kind (Vieta–Fibonacci polynomial) +evaluates on `2 * cos θ` to the value `sin ((n + 1) * θ) / sin θ`. -/ +@[simp] +theorem S_two_mul_complex_cos (n : ℤ) : (S ℂ n).eval (2 * cos θ) * sin θ = sin ((n + 1) * θ) := by + simp [S_eq_U_comp_half_mul_X] + end Complex /-! ### Real versions -/ @@ -115,6 +127,18 @@ value `sin ((n + 1) * θ) / sin θ`. -/ theorem U_real_cos : (U ℝ n).eval (cos θ) * sin θ = sin ((n + 1) * θ) := mod_cast U_complex_cos θ n +/-- The `n`-th rescaled Chebyshev polynomial of the first kind (Vieta–Lucas polynomial) evaluates on +`2 * cos θ` to the value `2 * cos (n * θ)`. -/ +@[simp] +theorem C_two_mul_real_cos : (C ℝ n).eval (2 * cos θ) = 2 * cos (n * θ) := by + simp [C_eq_two_mul_T_comp_half_mul_X] + +/-- The `n`-th rescaled Chebyshev polynomial of the second kind (Vieta–Fibonacci polynomial) +evaluates on `2 * cos θ` to the value `sin ((n + 1) * θ) / sin θ`. -/ +@[simp] +theorem S_two_mul_real_cos : (S ℝ n).eval (2 * cos θ) * sin θ = sin ((n + 1) * θ) := by + simp [S_eq_U_comp_half_mul_X] + end Real end Polynomial.Chebyshev diff --git a/Mathlib/RingTheory/Polynomial/Chebyshev.lean b/Mathlib/RingTheory/Polynomial/Chebyshev.lean index 4e9eb5a7fe9d2..3bed452ffb6b3 100644 --- a/Mathlib/RingTheory/Polynomial/Chebyshev.lean +++ b/Mathlib/RingTheory/Polynomial/Chebyshev.lean @@ -16,16 +16,21 @@ with integral coefficients. * `Polynomial.Chebyshev.T`: the Chebyshev polynomials of the first kind. * `Polynomial.Chebyshev.U`: the Chebyshev polynomials of the second kind. +* `Polynomial.Chebyshev.C`: the rescaled Chebyshev polynomials of the first kind (also known as the + Vieta–Lucas polynomials), given by $C_n(2x) = 2T_n(x)$. +* `Polynomial.Chebyshev.S`: the rescaled Chebyshev polynomials of the second kind (also known as the + Vieta–Fibonacci polynomials), given by $S_n(2x) = U_n(x)$. ## Main statements * The formal derivative of the Chebyshev polynomials of the first kind is a scalar multiple of the Chebyshev polynomials of the second kind. -* `Polynomial.Chebyshev.mul_T`, twice the product of the `m`-th and `k`-th Chebyshev polynomials of - the first kind is the sum of the `m + k`-th and `m - k`-th Chebyshev polynomials of the first - kind. +* `Polynomial.Chebyshev.T_mul_T`, twice the product of the `m`-th and `k`-th Chebyshev polynomials + of the first kind is the sum of the `m + k`-th and `m - k`-th Chebyshev polynomials of the first + kind. There is a similar statement `Polynomial.Chebyshev.C_mul_C` for the `C` polynomials. * `Polynomial.Chebyshev.T_mul`, the `(m * n)`-th Chebyshev polynomial of the first kind is the - composition of the `m`-th and `n`-th Chebyshev polynomials of the first kind. + composition of the `m`-th and `n`-th Chebyshev polynomials of the first kind. There is a similar + statement `Polynomial.Chebyshev.C_mul` for the `C` polynomials. ## Implementation details @@ -53,7 +58,7 @@ namespace Polynomial.Chebyshev open Polynomial -variable (R S : Type*) [CommRing R] [CommRing S] +variable (R R' : Type*) [CommRing R] [CommRing R'] /-- `T n` is the `n`-th Chebyshev polynomial of the first kind. -/ -- Well-founded definitions are now irreducible by default; @@ -100,7 +105,7 @@ theorem T_zero : T R 0 = 1 := rfl @[simp] theorem T_one : T R 1 = X := rfl -theorem T_neg_one : T R (-1) = X := (by ring : 2 * X * 1 - X = X) +theorem T_neg_one : T R (-1) = X := show 2 * X * 1 - X = X by ring theorem T_two : T R 2 = 2 * X ^ 2 - 1 := by simpa [pow_two, mul_assoc] using T_add_two R 0 @@ -217,10 +222,185 @@ theorem one_sub_X_sq_mul_U_eq_pol_in_T (n : ℤ) : (1 - X ^ 2) * U R n = X * T R (n + 1) - T R (n + 2) := by linear_combination T_eq_X_mul_T_sub_pol_U R n -variable {R S} +/-- `C n` is the `n`th rescaled Chebyshev polynomial of the first kind (also known as a Vieta–Lucas +polynomial), given by $C_n(2x) = 2T_n(x)$. See `Polynomial.Chebyshev.C_comp_two_mul_X`. -/ +@[semireducible] noncomputable def C : ℤ → R[X] + | 0 => 2 + | 1 => X + | (n : ℕ) + 2 => X * C (n + 1) - C n + | -((n : ℕ) + 1) => X * C (-n) - C (-n + 1) + termination_by n => Int.natAbs n + Int.natAbs (n - 1) + +@[simp] +theorem C_add_two : ∀ n, C R (n + 2) = X * C R (n + 1) - C R n + | (k : ℕ) => C.eq_3 R k + | -(k + 1 : ℕ) => by linear_combination (norm := (simp [Int.negSucc_eq]; ring_nf)) C.eq_4 R k + +theorem C_add_one (n : ℤ) : C R (n + 1) = X * C R n - C R (n - 1) := by + linear_combination (norm := ring_nf) C_add_two R (n - 1) + +theorem C_sub_two (n : ℤ) : C R (n - 2) = X * C R (n - 1) - C R n := by + linear_combination (norm := ring_nf) C_add_two R (n - 2) + +theorem C_sub_one (n : ℤ) : C R (n - 1) = X * C R n - C R (n + 1) := by + linear_combination (norm := ring_nf) C_add_two R (n - 1) + +theorem C_eq (n : ℤ) : C R n = X * C R (n - 1) - C R (n - 2) := by + linear_combination (norm := ring_nf) C_add_two R (n - 2) + +@[simp] +theorem C_zero : C R 0 = 2 := rfl + +@[simp] +theorem C_one : C R 1 = X := rfl + +theorem C_neg_one : C R (-1) = X := show X * 2 - X = X by ring + +theorem C_two : C R 2 = X ^ 2 - 2 := by + simpa [pow_two, mul_assoc] using C_add_two R 0 + +@[simp] +theorem C_neg (n : ℤ) : C R (-n) = C R n := by + induction n using Polynomial.Chebyshev.induct with + | zero => rfl + | one => show X * 2 - X = X; ring + | add_two n ih1 ih2 => + have h₁ := C_add_two R n + have h₂ := C_sub_two R (-n) + linear_combination (norm := ring_nf) (X:R[X]) * ih1 - ih2 - h₁ + h₂ + | neg_add_one n ih1 ih2 => + have h₁ := C_add_one R n + have h₂ := C_sub_one R (-n) + linear_combination (norm := ring_nf) (X:R[X]) * ih1 - ih2 + h₁ - h₂ + +theorem C_natAbs (n : ℤ) : C R n.natAbs = C R n := by + obtain h | h := Int.natAbs_eq n <;> nth_rw 2 [h]; simp + +theorem C_neg_two : C R (-2) = X ^ 2 - 2 := by simp [C_two] + +theorem C_comp_two_mul_X (n : ℤ) : (C R n).comp (2 * X) = 2 * T R n := by + induction n using Polynomial.Chebyshev.induct with + | zero => simp + | one => simp + | add_two n ih1 ih2 => + simp_rw [C_add_two, T_add_two, sub_comp, mul_comp, X_comp, ih1, ih2] + ring + | neg_add_one n ih1 ih2 => + simp_rw [C_sub_one, T_sub_one, sub_comp, mul_comp, X_comp, ih1, ih2] + ring + +theorem C_eq_two_mul_T_comp_half_mul_X [Invertible (2 : R)] (n : ℤ) : + C R n = 2 * (T R n).comp (Polynomial.C ⅟2 * X) := by + have := congr_arg (·.comp (Polynomial.C ⅟2 * X)) (C_comp_two_mul_X R n) + simp_rw [comp_assoc, mul_comp, ofNat_comp, X_comp, ← mul_assoc, ← C_eq_natCast, ← C_mul, + Nat.cast_ofNat, mul_invOf_self', map_one, one_mul, comp_X, map_ofNat] at this + assumption + +theorem T_eq_half_mul_C_comp_two_mul_X [Invertible (2 : R)] (n : ℤ) : + T R n = Polynomial.C ⅟2 * (C R n).comp (2 * X) := by + rw [C_comp_two_mul_X, ← mul_assoc, ← map_ofNat Polynomial.C 2, ← map_mul, invOf_mul_self', + map_one, one_mul] + +/-- `S n` is the `n`th rescaled Chebyshev polynomial of the second kind (also known as a +Vieta–Fibonacci polynomial), given by $S_n(2x) = U_n(x)$. See +`Polynomial.Chebyshev.S_comp_two_mul_X`. -/ +@[semireducible] noncomputable def S : ℤ → R[X] + | 0 => 1 + | 1 => X + | (n : ℕ) + 2 => X * S (n + 1) - S n + | -((n : ℕ) + 1) => X * S (-n) - S (-n + 1) + termination_by n => Int.natAbs n + Int.natAbs (n - 1) + +@[simp] +theorem S_add_two : ∀ n, S R (n + 2) = X * S R (n + 1) - S R n + | (k : ℕ) => S.eq_3 R k + | -(k + 1 : ℕ) => by linear_combination (norm := (simp [Int.negSucc_eq]; ring_nf)) S.eq_4 R k + +theorem S_add_one (n : ℤ) : S R (n + 1) = X * S R n - S R (n - 1) := by + linear_combination (norm := ring_nf) S_add_two R (n - 1) + +theorem S_sub_two (n : ℤ) : S R (n - 2) = X * S R (n - 1) - S R n := by + linear_combination (norm := ring_nf) S_add_two R (n - 2) + +theorem S_sub_one (n : ℤ) : S R (n - 1) = X * S R n - S R (n + 1) := by + linear_combination (norm := ring_nf) S_add_two R (n - 1) + +theorem S_eq (n : ℤ) : S R n = X * S R (n - 1) - S R (n - 2) := by + linear_combination (norm := ring_nf) S_add_two R (n - 2) + +@[simp] +theorem S_zero : S R 0 = 1 := rfl + +@[simp] +theorem S_one : S R 1 = X := rfl @[simp] -theorem map_T (f : R →+* S) (n : ℤ) : map f (T R n) = T S n := by +theorem S_neg_one : S R (-1) = 0 := by simpa using S_sub_one R 0 + +theorem S_two : S R 2 = X ^ 2 - 1 := by + have := S_add_two R 0 + simp only [zero_add, S_one, S_zero] at this + linear_combination this + +@[simp] +theorem S_neg_two : S R (-2) = -1 := by + simpa [zero_sub, Int.reduceNeg, S_neg_one, mul_zero, S_zero] using S_sub_two R 0 + +theorem S_neg_sub_one (n : ℤ) : S R (-n - 1) = -S R (n - 1) := by + induction n using Polynomial.Chebyshev.induct with + | zero => simp + | one => simp + | add_two n ih1 ih2 => + have h₁ := S_add_one R n + have h₂ := S_sub_two R (-n - 1) + linear_combination (norm := ring_nf) (X:R[X]) * ih1 - ih2 + h₁ + h₂ + | neg_add_one n ih1 ih2 => + have h₁ := S_eq R n + have h₂ := S_sub_two R (-n) + linear_combination (norm := ring_nf) (X:R[X]) * ih1 - ih2 + h₁ + h₂ + +theorem S_neg (n : ℤ) : S R (-n) = -S R (n - 2) := by simpa [sub_sub] using S_neg_sub_one R (n - 1) + +@[simp] +theorem S_neg_sub_two (n : ℤ) : S R (-n - 2) = -S R n := by + simpa [sub_eq_add_neg, add_comm] using S_neg R (n + 2) + +theorem S_comp_two_mul_X (n : ℤ) : (S R n).comp (2 * X) = U R n := by + induction n using Polynomial.Chebyshev.induct with + | zero => simp + | one => simp + | add_two n ih1 ih2 => simp_rw [U_add_two, S_add_two, sub_comp, mul_comp, X_comp, ih1, ih2] + | neg_add_one n ih1 ih2 => simp_rw [U_sub_one, S_sub_one, sub_comp, mul_comp, X_comp, ih1, ih2] + +theorem S_eq_U_comp_half_mul_X [Invertible (2 : R)] (n : ℤ) : + S R n = (U R n).comp (Polynomial.C ⅟2 * X) := by + have := congr_arg (·.comp (Polynomial.C ⅟2 * X)) (S_comp_two_mul_X R n) + simp_rw [comp_assoc, mul_comp, ofNat_comp, X_comp, ← mul_assoc, ← C_eq_natCast, ← C_mul, + Nat.cast_ofNat, mul_invOf_self', map_one, one_mul, comp_X] at this + assumption + +theorem S_eq_X_mul_S_add_C (n : ℤ) : 2 * S R (n + 1) = X * S R n + C R (n + 1) := by + induction n using Polynomial.Chebyshev.induct with + | zero => simp [two_mul] + | one => simp [S_two, C_two]; ring + | add_two n ih1 ih2 => + have h₁ := S_add_two R (n + 1) + have h₂ := S_add_two R n + have h₃ := C_add_two R (n + 1) + linear_combination (norm := ring_nf) -h₃ - (X:R[X]) * h₂ + 2 * h₁ + (X:R[X]) * ih1 - ih2 + | neg_add_one n ih1 ih2 => + have h₁ := S_add_two R (-n - 1) + have h₂ := S_add_two R (-n) + have h₃ := C_add_two R (-n) + linear_combination (norm := ring_nf) -h₃ + 2 * h₂ - (X:R[X]) * h₁ - ih2 + (X:R[X]) * ih1 + +theorem C_eq_S_sub_X_mul_S (n : ℤ) : C R n = 2 * S R n - X * S R (n - 1) := by + linear_combination (norm := ring_nf) - S_eq_X_mul_S_add_C R (n - 1) + +variable {R R'} + +@[simp] +theorem map_T (f : R →+* R') (n : ℤ) : map f (T R n) = T R' n := by induction n using Polynomial.Chebyshev.induct with | zero => simp | one => simp @@ -232,7 +412,7 @@ theorem map_T (f : R →+* S) (n : ℤ) : map f (T R n) = T S n := by ih2] @[simp] -theorem map_U (f : R →+* S) (n : ℤ) : map f (U R n) = U S n := by +theorem map_U (f : R →+* R') (n : ℤ) : map f (U R n) = U R' n := by induction n using Polynomial.Chebyshev.induct with | zero => simp | one => simp @@ -243,6 +423,26 @@ theorem map_U (f : R →+* S) (n : ℤ) : map f (U R n) = U S n := by simp_rw [U_sub_one, Polynomial.map_sub, Polynomial.map_mul, Polynomial.map_ofNat, map_X, ih1, ih2] +@[simp] +theorem map_C (f : R →+* R') (n : ℤ) : map f (C R n) = C R' n := by + induction n using Polynomial.Chebyshev.induct with + | zero => simp + | one => simp + | add_two n ih1 ih2 => + simp_rw [C_add_two, Polynomial.map_sub, Polynomial.map_mul, map_X, ih1, ih2] + | neg_add_one n ih1 ih2 => + simp_rw [C_sub_one, Polynomial.map_sub, Polynomial.map_mul, map_X, ih1, ih2] + +@[simp] +theorem map_S (f : R →+* R') (n : ℤ) : map f (S R n) = S R' n := by + induction n using Polynomial.Chebyshev.induct with + | zero => simp + | one => simp + | add_two n ih1 ih2 => + simp_rw [S_add_two, Polynomial.map_sub, Polynomial.map_mul, map_X, ih1, ih2] + | neg_add_one n ih1 ih2 => + simp_rw [S_sub_one, Polynomial.map_sub, Polynomial.map_mul, map_X, ih1, ih2] + theorem T_derivative_eq_U (n : ℤ) : derivative (T R n) = n * U R (n - 1) := by induction n using Polynomial.Chebyshev.induct with | zero => simp @@ -284,7 +484,7 @@ variable (R) /-- Twice the product of two Chebyshev `T` polynomials is the sum of two other Chebyshev `T` polynomials. -/ -theorem mul_T (m k : ℤ) : 2 * T R m * T R k = T R (m + k) + T R (m - k) := by +theorem T_mul_T (m k : ℤ) : 2 * T R m * T R k = T R (m + k) + T R (m - k) := by induction k using Polynomial.Chebyshev.induct with | zero => simp [two_mul] | one => rw [T_add_one, T_one]; ring @@ -299,20 +499,55 @@ theorem mul_T (m k : ℤ) : 2 * T R m * T R k = T R (m + k) + T R (m - k) := by have h₃ := T_add_two R (-k - 1) linear_combination (norm := ring_nf) 2 * T R m * h₃ - h₂ - h₁ - ih2 + 2 * (X : R[X]) * ih1 +@[deprecated (since := "2024-12-03")] alias mul_T := T_mul_T + +/-- The product of two Chebyshev `C` polynomials is the sum of two other Chebyshev `C` polynomials. +-/ +theorem C_mul_C (m k : ℤ) : C R m * C R k = C R (m + k) + C R (m - k) := by + induction k using Polynomial.Chebyshev.induct with + | zero => simp [mul_two] + | one => rw [C_add_one, C_one]; ring + | add_two k ih1 ih2 => + have h₁ := C_add_two R (m + k) + have h₂ := C_sub_two R (m - k) + have h₃ := C_add_two R k + linear_combination (norm := ring_nf) C R m * h₃ - h₂ - h₁ - ih2 + (X:R[X]) * ih1 + | neg_add_one k ih1 ih2 => + have h₁ := C_add_two R (m + (-k - 1)) + have h₂ := C_sub_two R (m - (-k - 1)) + have h₃ := C_add_two R (-k - 1) + linear_combination (norm := ring_nf) C R m * h₃ - h₂ - h₁ - ih2 + (X:R[X]) * ih1 + /-- The `(m * n)`-th Chebyshev `T` polynomial is the composition of the `m`-th and `n`-th. -/ theorem T_mul (m n : ℤ) : T R (m * n) = (T R m).comp (T R n) := by induction m using Polynomial.Chebyshev.induct with | zero => simp | one => simp | add_two m ih1 ih2 => - have h₁ := mul_T R ((m + 1) * n) n + have h₁ := T_mul_T R ((m + 1) * n) n have h₂ := congr_arg (comp · (T R n)) <| T_add_two R m simp only [sub_comp, mul_comp, ofNat_comp, X_comp] at h₂ linear_combination (norm := ring_nf) -ih2 - h₂ - h₁ + 2 * T R n * ih1 | neg_add_one m ih1 ih2 => - have h₁ := mul_T R ((-m) * n) n + have h₁ := T_mul_T R ((-m) * n) n have h₂ := congr_arg (comp · (T R n)) <| T_add_two R (-m - 1) simp only [sub_comp, mul_comp, ofNat_comp, X_comp] at h₂ linear_combination (norm := ring_nf) -ih2 - h₂ - h₁ + 2 * T R n * ih1 +/-- The `(m * n)`-th Chebyshev `C` polynomial is the composition of the `m`-th and `n`-th. -/ +theorem C_mul (m n : ℤ) : C R (m * n) = (C R m).comp (C R n) := by + induction m using Polynomial.Chebyshev.induct with + | zero => simp + | one => simp + | add_two m ih1 ih2 => + have h₁ := C_mul_C R ((m + 1) * n) n + have h₂ := congr_arg (comp · (C R n)) <| C_add_two R m + simp only [sub_comp, mul_comp, X_comp] at h₂ + linear_combination (norm := ring_nf) -ih2 - h₂ - h₁ + C R n * ih1 + | neg_add_one m ih1 ih2 => + have h₁ := C_mul_C R ((-m) * n) n + have h₂ := congr_arg (comp · (C R n)) <| C_add_two R (-m - 1) + simp only [sub_comp, mul_comp, X_comp] at h₂ + linear_combination (norm := ring_nf) -ih2 - h₂ - h₁ + C R n * ih1 + end Polynomial.Chebyshev diff --git a/Mathlib/RingTheory/Polynomial/Dickson.lean b/Mathlib/RingTheory/Polynomial/Dickson.lean index 1872e6ac8a9e2..b563e9a572c2a 100644 --- a/Mathlib/RingTheory/Polynomial/Dickson.lean +++ b/Mathlib/RingTheory/Polynomial/Dickson.lean @@ -135,25 +135,45 @@ private theorem two_mul_C_half_eq_one [Invertible (2 : R)] : 2 * C (⅟ 2 : R) = private theorem C_half_mul_two_eq_one [Invertible (2 : R)] : C (⅟ 2 : R) * 2 = 1 := by rw [mul_comm, two_mul_C_half_eq_one] -theorem dickson_one_one_eq_chebyshev_T [Invertible (2 : R)] : - ∀ n, dickson 1 (1 : R) n = 2 * (Chebyshev.T R n).comp (C (⅟ 2) * X) +theorem dickson_one_one_eq_chebyshev_C : ∀ n, dickson 1 (1 : R) n = Chebyshev.C R n | 0 => by - simp only [Chebyshev.T_zero, mul_one, one_comp, dickson_zero] + simp only [Chebyshev.C_zero, mul_one, one_comp, dickson_zero] norm_num | 1 => by - rw [dickson_one, Nat.cast_one, Chebyshev.T_one, X_comp, ← mul_assoc, two_mul_C_half_eq_one, - one_mul] + rw [dickson_one, Nat.cast_one, Chebyshev.C_one] | n + 2 => by - rw [dickson_add_two, C_1, Nat.cast_add, Nat.cast_two, Chebyshev.T_add_two, - dickson_one_one_eq_chebyshev_T (n + 1), dickson_one_one_eq_chebyshev_T n, sub_comp, mul_comp, - mul_comp, X_comp, ofNat_comp] - simp_rw [← mul_assoc, Nat.cast_ofNat, two_mul_C_half_eq_one, Nat.cast_add, Nat.cast_one] + rw [dickson_add_two, C_1, Nat.cast_add, Nat.cast_two, Chebyshev.C_add_two, + dickson_one_one_eq_chebyshev_C (n + 1), dickson_one_one_eq_chebyshev_C n] + push_cast ring +theorem dickson_one_one_eq_chebyshev_T [Invertible (2 : R)] (n : ℕ) : + dickson 1 (1 : R) n = 2 * (Chebyshev.T R n).comp (C (⅟ 2) * X) := + (dickson_one_one_eq_chebyshev_C R n).trans (Chebyshev.C_eq_two_mul_T_comp_half_mul_X R n) + theorem chebyshev_T_eq_dickson_one_one [Invertible (2 : R)] (n : ℕ) : - Chebyshev.T R n = C (⅟ 2) * (dickson 1 1 n).comp (2 * X) := by - rw [dickson_one_one_eq_chebyshev_T, mul_comp, ofNat_comp, comp_assoc, mul_comp, C_comp, X_comp] - simp_rw [← mul_assoc, Nat.cast_ofNat, C_half_mul_two_eq_one, one_mul, comp_X] + Chebyshev.T R n = C (⅟ 2) * (dickson 1 1 n).comp (2 * X) := + dickson_one_one_eq_chebyshev_C R n ▸ Chebyshev.T_eq_half_mul_C_comp_two_mul_X R n + +theorem dickson_two_one_eq_chebyshev_S : ∀ n, dickson 2 (1 : R) n = Chebyshev.S R n + | 0 => by + simp only [Chebyshev.S_zero, mul_one, one_comp, dickson_zero] + norm_num + | 1 => by + rw [dickson_one, Nat.cast_one, Chebyshev.S_one] + | n + 2 => by + rw [dickson_add_two, C_1, Nat.cast_add, Nat.cast_two, Chebyshev.S_add_two, + dickson_two_one_eq_chebyshev_S (n + 1), dickson_two_one_eq_chebyshev_S n] + push_cast + ring + +theorem dickson_two_one_eq_chebyshev_U [Invertible (2 : R)] (n : ℕ) : + dickson 2 (1 : R) n = (Chebyshev.U R n).comp (C (⅟ 2) * X) := + (dickson_two_one_eq_chebyshev_S R n).trans (Chebyshev.S_eq_U_comp_half_mul_X R n) + +theorem chebyshev_U_eq_dickson_two_one (n : ℕ) : + Chebyshev.U R n = (dickson 2 (1 : R) n).comp (2 * X) := + dickson_two_one_eq_chebyshev_S R n ▸ (Chebyshev.S_comp_two_mul_X R n).symm /-- The `(m * n)`-th Dickson polynomial of the first kind is the composition of the `m`-th and `n`-th. -/ From ea515836ea82a84277b33b28c840a24da4318c4f Mon Sep 17 00:00:00 2001 From: Antoine Chambert-Loir Date: Wed, 4 Dec 2024 05:57:33 +0000 Subject: [PATCH 504/829] feat(RingTheory/MvPolynomial/Groebner): division algorithm wrt monomial orders (#19454) Formalize the division algorithm with respect to a monomial order This is part of an ongoing work to formalize Groebner theory. For the moment, only the lexicographic order is present in mathlib, but subsequent PRs will introduce the homogeneous lexicographic order and the homogeneous reverse lexicographic order. --- Mathlib.lean | 2 + Mathlib/Data/Finset/Lattice/Fold.lean | 16 + Mathlib/Data/Finsupp/MonomialOrder.lean | 1 - Mathlib/RingTheory/MvPolynomial/Groebner.lean | 238 +++++++++++ .../MvPolynomial/MonomialOrder.lean | 388 ++++++++++++++++++ 5 files changed, 644 insertions(+), 1 deletion(-) create mode 100644 Mathlib/RingTheory/MvPolynomial/Groebner.lean create mode 100644 Mathlib/RingTheory/MvPolynomial/MonomialOrder.lean diff --git a/Mathlib.lean b/Mathlib.lean index b24ff15bb5cab..144457d89beae 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4433,9 +4433,11 @@ import Mathlib.RingTheory.MvPolynomial import Mathlib.RingTheory.MvPolynomial.Basic import Mathlib.RingTheory.MvPolynomial.EulerIdentity import Mathlib.RingTheory.MvPolynomial.FreeCommRing +import Mathlib.RingTheory.MvPolynomial.Groebner import Mathlib.RingTheory.MvPolynomial.Homogeneous import Mathlib.RingTheory.MvPolynomial.Ideal import Mathlib.RingTheory.MvPolynomial.Localization +import Mathlib.RingTheory.MvPolynomial.MonomialOrder import Mathlib.RingTheory.MvPolynomial.Symmetric.Defs import Mathlib.RingTheory.MvPolynomial.Symmetric.FundamentalTheorem import Mathlib.RingTheory.MvPolynomial.Symmetric.NewtonIdentities diff --git a/Mathlib/Data/Finset/Lattice/Fold.lean b/Mathlib/Data/Finset/Lattice/Fold.lean index a9cc52080c652..97a141e5d4d22 100644 --- a/Mathlib/Data/Finset/Lattice/Fold.lean +++ b/Mathlib/Data/Finset/Lattice/Fold.lean @@ -634,6 +634,22 @@ protected theorem sup_lt_iff (ha : ⊥ < a) : s.sup f < a ↔ ∀ b ∈ s, f b < Finset.cons_induction_on s (fun _ => ha) fun c t hc => by simpa only [sup_cons, sup_lt_iff, mem_cons, forall_eq_or_imp] using And.imp_right⟩ +theorem sup_mem_of_nonempty (hs : s.Nonempty) : s.sup f ∈ f '' s := by + classical + induction s using Finset.induction with + | empty => exfalso; simp only [Finset.not_nonempty_empty] at hs + | @insert a s _ h => + rw [Finset.sup_insert (b := a) (s := s) (f := f)] + by_cases hs : s = ∅ + · simp [hs] + · rw [← ne_eq, ← Finset.nonempty_iff_ne_empty] at hs + simp only [Finset.coe_insert] + rcases le_total (f a) (s.sup f) with (ha | ha) + · rw [sup_eq_right.mpr ha] + exact Set.image_mono (Set.subset_insert a s) (h hs) + · rw [sup_eq_left.mpr ha] + apply Set.mem_image_of_mem _ (Set.mem_insert a ↑s) + end OrderBot section OrderTop diff --git a/Mathlib/Data/Finsupp/MonomialOrder.lean b/Mathlib/Data/Finsupp/MonomialOrder.lean index 431fcb7d3ce34..788df94797c35 100644 --- a/Mathlib/Data/Finsupp/MonomialOrder.lean +++ b/Mathlib/Data/Finsupp/MonomialOrder.lean @@ -135,7 +135,6 @@ example : toLex (Finsupp.single 1 1) < toLex (Finsupp.single 0 1) := by example : toLex (Finsupp.single 1 1) < toLex (Finsupp.single 0 2) := by use 0; simp - variable {σ : Type*} [LinearOrder σ] /-- The lexicographic order on `σ →₀ ℕ`, as a `MonomialOrder` -/ diff --git a/Mathlib/RingTheory/MvPolynomial/Groebner.lean b/Mathlib/RingTheory/MvPolynomial/Groebner.lean new file mode 100644 index 0000000000000..a2149146b532f --- /dev/null +++ b/Mathlib/RingTheory/MvPolynomial/Groebner.lean @@ -0,0 +1,238 @@ +/- +Copyright (c) 2024 Antoine Chambert-Loir. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Antoine Chambert-Loir +-/ +import Mathlib.Data.Finsupp.Lex +import Mathlib.Data.Finsupp.MonomialOrder +import Mathlib.Data.Finsupp.WellFounded +import Mathlib.Data.List.TFAE +import Mathlib.RingTheory.MvPolynomial.Homogeneous +import Mathlib.RingTheory.MvPolynomial.MonomialOrder + +/-! # Division algorithm with respect to monomial orders + +We provide a division algorithm with respect to monomial orders in polynomial rings. +Let `R` be a commutative ring, `σ` a type of indeterminates and `m : MonomialOrder σ` +a monomial ordering on `σ →₀ ℕ`. + +Consider a family of polynomials `b : ι → MvPolynomial σ R` with invertible leading coefficients +(with respect to `m`) : we assume `hb : ∀ i, IsUnit (m.lCoeff (b i))`). + +* `MonomialOrder.div hb f` furnishes + - a finitely supported family `g : ι →₀ MvPolynomial σ R` + - and a “remainder” `r : MvPolynomial σ R` +such that the three properties hold: + (1) One has `f = ∑ (g i) * (b i) + r` + (2) For every `i`, `m.degree ((g i) * (b i)` is less than or equal to that of `f` + (3) For every `i`, every monomial in the support of `r` is strictly smaller + than the leading term of `b i`, + +The proof is done by induction, using two standard constructions + +* `MonomialOrder.subLTerm f` deletes the leading term of a polynomial `f` + +* `MonomialOrder.reduce hb f` subtracts from `f` the appropriate multiple of `b : MvPolynomial σ R`, +provided `IsUnit (m.lCoeff b)`. + +* `MonomialOrder.div_set` is the variant of `MonomialOrder.div` for a set of polynomials. + +## Reference : [Becker-Weispfenning1993] + +## TODO + +* Prove that under `Field F`, `IsUnit (m.lCoeff (b i))` is equivalent to `b i ≠ 0`. + +-/ + +namespace MonomialOrder + +open MvPolynomial + +open scoped MonomialOrder + +variable {σ : Type*} {m : MonomialOrder σ} {R : Type*} [CommRing R] + +variable (m) in +/-- Delete the leading term in a multivariate polynomial (for some monomial order) -/ +noncomputable def subLTerm (f : MvPolynomial σ R) : MvPolynomial σ R := + f - monomial (m.degree f) (m.lCoeff f) + +theorem degree_sub_LTerm_le (f : MvPolynomial σ R) : + m.degree (m.subLTerm f) ≼[m] m.degree f := by + apply le_trans degree_sub_le + simp only [sup_le_iff, le_refl, true_and] + apply degree_monomial_le + +theorem degree_sub_LTerm_lt {f : MvPolynomial σ R} (hf : m.degree f ≠ 0) : + m.degree (m.subLTerm f) ≺[m] m.degree f := by + rw [lt_iff_le_and_ne] + refine ⟨degree_sub_LTerm_le f, ?_⟩ + classical + intro hf' + simp only [EmbeddingLike.apply_eq_iff_eq] at hf' + have : m.subLTerm f ≠ 0 := by + intro h + simp only [h, degree_zero] at hf' + exact hf hf'.symm + rw [← coeff_degree_ne_zero_iff (m := m), hf'] at this + apply this + simp [subLTerm, coeff_monomial, lCoeff] + +variable (m) in +/-- Reduce a polynomial modulo a polynomial with unit leading term (for some monomial order) -/ +noncomputable def reduce {b : MvPolynomial σ R} (hb : IsUnit (m.lCoeff b)) (f : MvPolynomial σ R) : + MvPolynomial σ R := + f - monomial (m.degree f - m.degree b) (hb.unit⁻¹ * m.lCoeff f) * b + +theorem degree_reduce_lt {f b : MvPolynomial σ R} (hb : IsUnit (m.lCoeff b)) + (hbf : m.degree b ≤ m.degree f) (hf : m.degree f ≠ 0) : + m.degree (m.reduce hb f) ≺[m] m.degree f := by + have H : m.degree f = + m.degree ((monomial (m.degree f - m.degree b)) (hb.unit⁻¹ * m.lCoeff f)) + + m.degree b := by + classical + rw [degree_monomial, if_neg] + · ext d + rw [tsub_add_cancel_of_le hbf] + · simp only [Units.mul_right_eq_zero, lCoeff_eq_zero_iff] + intro hf0 + apply hf + simp [hf0] + have H' : coeff (m.degree f) (m.reduce hb f) = 0 := by + simp only [reduce, coeff_sub, sub_eq_zero] + nth_rewrite 2 [H] + rw [coeff_mul_of_degree_add (m := m), lCoeff_monomial] + rw [mul_comm, ← mul_assoc] + simp only [IsUnit.mul_val_inv, one_mul] + rfl + rw [lt_iff_le_and_ne] + constructor + · classical + apply le_trans degree_sub_le + simp only [sup_le_iff, le_refl, true_and] + apply le_of_le_of_eq degree_mul_le + rw [m.toSyn.injective.eq_iff] + exact H.symm + · intro K + simp only [EmbeddingLike.apply_eq_iff_eq] at K + nth_rewrite 1 [← K] at H' + change lCoeff m _ = 0 at H' + rw [lCoeff_eq_zero_iff] at H' + rw [H', degree_zero] at K + exact hf K.symm + +theorem div {ι : Type*} {b : ι → MvPolynomial σ R} + (hb : ∀ i, IsUnit (m.lCoeff (b i))) (f : MvPolynomial σ R) : + ∃ (g : ι →₀ (MvPolynomial σ R)) (r : MvPolynomial σ R), + f = Finsupp.linearCombination _ b g + r ∧ + (∀ i, m.degree (b i * (g i)) ≼[m] m.degree f) ∧ + (∀ c ∈ r.support, ∀ i, ¬ (m.degree (b i) ≤ c)) := by + by_cases hb' : ∃ i, m.degree (b i) = 0 + · obtain ⟨i, hb0⟩ := hb' + use Finsupp.single i ((hb i).unit⁻¹ • f), 0 + constructor + · simp only [Finsupp.linearCombination_single, smul_eq_mul, add_zero] + simp only [smul_mul_assoc, ← smul_eq_iff_eq_inv_smul, Units.smul_isUnit] + nth_rewrite 2 [eq_C_of_degree_eq_zero hb0] + rw [mul_comm, smul_eq_C_mul] + constructor + · intro j + by_cases hj : j = i + · apply le_trans degree_mul_le + simp only [hj, hb0, Finsupp.single_eq_same, zero_add] + apply le_of_eq + simp only [EmbeddingLike.apply_eq_iff_eq] + apply degree_smul (Units.isRegular _) + · simp only [Finsupp.single_eq_of_ne (Ne.symm hj), mul_zero, degree_zero, map_zero] + apply bot_le + · simp + push_neg at hb' + by_cases hf0 : f = 0 + · refine ⟨0, 0, by simp [hf0], ?_, by simp⟩ + intro b + simp only [Finsupp.coe_zero, Pi.zero_apply, mul_zero, degree_zero, map_zero] + exact bot_le + by_cases hf : ∃ i, m.degree (b i) ≤ m.degree f + · obtain ⟨i, hf⟩ := hf + have deg_reduce : m.degree (m.reduce (hb i) f) ≺[m] m.degree f := by + apply degree_reduce_lt (hb i) hf + intro hf0' + apply hb' i + simpa [hf0'] using hf + obtain ⟨g', r', H'⟩ := div hb (m.reduce (hb i) f) + use g' + + Finsupp.single i (monomial (m.degree f - m.degree (b i)) ((hb i).unit⁻¹ * m.lCoeff f)) + use r' + constructor + · rw [map_add, add_assoc, add_comm _ r', ← add_assoc, ← H'.1] + simp [reduce] + constructor + · rintro j + simp only [Finsupp.coe_add, Pi.add_apply] + rw [mul_add] + apply le_trans degree_add_le + simp only [sup_le_iff] + constructor + · exact le_trans (H'.2.1 _) (le_of_lt deg_reduce) + · classical + rw [Finsupp.single_apply] + split_ifs with hc + · apply le_trans degree_mul_le + simp only [map_add] + apply le_of_le_of_eq (add_le_add_left (degree_monomial_le _) _) + simp only [← hc] + rw [← map_add, m.toSyn.injective.eq_iff] + rw [add_tsub_cancel_of_le] + exact hf + · simp only [mul_zero, degree_zero, map_zero] + exact bot_le + · exact H'.2.2 + · push_neg at hf + suffices ∃ (g' : ι →₀ MvPolynomial σ R), ∃ r', + (m.subLTerm f = Finsupp.linearCombination (MvPolynomial σ R) b g' + r') ∧ + (∀ i, m.degree ((b i) * (g' i)) ≼[m] m.degree (m.subLTerm f)) ∧ + (∀ c ∈ r'.support, ∀ i, ¬ m.degree (b i) ≤ c) by + obtain ⟨g', r', H'⟩ := this + use g', r' + monomial (m.degree f) (m.lCoeff f) + constructor + · simp [← add_assoc, ← H'.1, subLTerm] + constructor + · exact fun b ↦ le_trans (H'.2.1 b) (degree_sub_LTerm_le f) + · intro c hc i + by_cases hc' : c ∈ r'.support + · exact H'.2.2 c hc' i + · convert hf i + classical + have := MvPolynomial.support_add hc + rw [Finset.mem_union, Classical.or_iff_not_imp_left] at this + simpa only [Finset.mem_singleton] using support_monomial_subset (this hc') + by_cases hf'0 : m.subLTerm f = 0 + · refine ⟨0, 0, by simp [hf'0], ?_, by simp⟩ + intro b + simp only [Finsupp.coe_zero, Pi.zero_apply, mul_zero, degree_zero, map_zero] + exact bot_le + · exact (div hb) (m.subLTerm f) +termination_by WellFounded.wrap + ((isWellFounded_iff m.syn fun x x_1 ↦ x < x_1).mp m.wf) (m.toSyn (m.degree f)) +decreasing_by +· exact deg_reduce +· apply degree_sub_LTerm_lt + intro hf0 + apply hf'0 + simp only [subLTerm, sub_eq_zero] + nth_rewrite 1 [eq_C_of_degree_eq_zero hf0, hf0] + simp + +theorem div_set {B : Set (MvPolynomial σ R)} + (hB : ∀ b ∈ B, IsUnit (m.lCoeff b)) (f : MvPolynomial σ R) : + ∃ (g : B →₀ (MvPolynomial σ R)) (r : MvPolynomial σ R), + f = Finsupp.linearCombination _ (fun (b : B) ↦ (b : MvPolynomial σ R)) g + r ∧ + (∀ (b : B), m.degree ((b : MvPolynomial σ R) * (g b)) ≼[m] m.degree f) ∧ + (∀ c ∈ r.support, ∀ b ∈ B, ¬ (m.degree b ≤ c)) := by + obtain ⟨g, r, H⟩ := m.div (b := fun (p : B) ↦ p) (fun b ↦ hB b b.prop) f + exact ⟨g, r, H.1, H.2.1, fun c hc b hb ↦ H.2.2 c hc ⟨b, hb⟩⟩ + +end MonomialOrder + + diff --git a/Mathlib/RingTheory/MvPolynomial/MonomialOrder.lean b/Mathlib/RingTheory/MvPolynomial/MonomialOrder.lean new file mode 100644 index 0000000000000..674a8ee44cb9d --- /dev/null +++ b/Mathlib/RingTheory/MvPolynomial/MonomialOrder.lean @@ -0,0 +1,388 @@ +/- +Copyright (c) 2024 Antoine Chambert-Loir. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Antoine Chambert-Loir +-/ +import Mathlib.RingTheory.MvPolynomial.Homogeneous +import Mathlib.Data.Finsupp.Lex +import Mathlib.Data.Finsupp.WellFounded +import Mathlib.Data.List.TFAE +import Mathlib.Data.Finsupp.MonomialOrder + +/-! # Degree and leading coefficient of polynomials with respect to a monomial order + +We consider a type `σ` of indeterminates and a commutative semiring `R` +and a monomial order `m : MonomialOrder σ`. + +* `m.degree f` is the degree of `f` for the monomial ordering `m` + +* `m.lCoeff f` is the leading coefficient of `f` for the monomial ordering `m` + +* `m.lCoeff_ne_zero_iff f` asserts that this coefficient is nonzero iff `f ≠ 0`. + +* in a field, `m.lCoeff_is_unit_iff f` asserts that this coefficient is a unit iff `f ≠ 0`. + +* `m.degree_add_le` : the `m.degree` of `f + g` is smaller than or equal to the supremum +of those of `f` and `g` + +* `m.degree_add_of_lt h` : the `m.degree` of `f + g` is equal to that of `f` +if the `m.degree` of `g` is strictly smaller than that `f` + +* `m.lCoeff_add_of_lt h`: then, the leading coefficient of `f + g` is that of `f` . + +* `m.degree_add_of_ne h` : the `m.degree` of `f + g` is equal to that the supremum +of those of `f` and `g` if they are distinct + +* `m.degree_sub_le` : the `m.degree` of `f - g` is smaller than or equal to the supremum +of those of `f` and `g` + +* `m.degree_sub_of_lt h` : the `m.degree` of `f - g` is equal to that of `f` +if the `m.degree` of `g` is strictly smaller than that `f` + +* `m.lCoeff_sub_of_lt h`: then, the leading coefficient of `f - g` is that of `f` . + +* `m.degree_mul_le`: the `m.degree` of `f * g` is smaller than or equal to the sum of those of +`f` and `g`. + +* `m.degree_mul_of_isRegular_left`, `m.degree_mul_of_isRegular_right` and `m.degree_mul` +assert the equality when the leading coefficient of `f` or `g` is regular, +or when `R` is a domain and `f` and `g` are nonzero. + +* `m.lCoeff_mul_of_isRegular_left`, `m.lCoeff_mul_of_isRegular_right` and `m.lCoeff_mul` +say that `m.lCoeff (f * g) = m.lCoeff f * m.lCoeff g` + +## Reference + +[Becker-Weispfenning1993] + +-/ + +namespace MonomialOrder + +open MvPolynomial + +open scoped MonomialOrder + +variable {σ : Type*} {m : MonomialOrder σ} + +section Semiring + +variable {R : Type*} [CommSemiring R] + +variable (m) in +/-- the degree of a multivariate polynomial with respect to a monomial ordering -/ +def degree {R : Type*} [CommSemiring R] (f : MvPolynomial σ R) : σ →₀ ℕ := + m.toSyn.symm (f.support.sup m.toSyn) + +variable (m) in +/-- the leading coefficient of a multivariate polynomial with respect to a monomial ordering -/ +def lCoeff {R : Type*} [CommSemiring R] (f : MvPolynomial σ R) : R := + f.coeff (m.degree f) + +@[simp] +theorem degree_zero : m.degree (0 : MvPolynomial σ R) = 0 := by + simp [degree] + +@[simp] +theorem lCoeff_zero : m.lCoeff (0 : MvPolynomial σ R) = 0 := by + simp [degree, lCoeff] + +theorem degree_monomial_le {d : σ →₀ ℕ} (c : R) : + m.degree (monomial d c) ≼[m] d := by + simp only [degree, AddEquiv.apply_symm_apply] + apply le_trans (Finset.sup_mono support_monomial_subset) + simp only [Finset.sup_singleton, le_refl] + +theorem degree_monomial {d : σ →₀ ℕ} (c : R) [Decidable (c = 0)] : + m.degree (monomial d c) = if c = 0 then 0 else d := by + simp only [degree, support_monomial] + split_ifs with hc <;> simp + +@[simp] +theorem lCoeff_monomial {d : σ →₀ ℕ} (c : R) : + m.lCoeff (monomial d c) = c := by + classical + simp only [lCoeff, degree_monomial] + split_ifs with hc <;> simp [hc] + +theorem degree_le_iff {f : MvPolynomial σ R} {d : σ →₀ ℕ} : + m.degree f ≼[m] d ↔ ∀ c ∈ f.support, c ≼[m] d := by + unfold degree + simp only [AddEquiv.apply_symm_apply, Finset.sup_le_iff, mem_support_iff, ne_eq] + +theorem degree_lt_iff {f : MvPolynomial σ R} {d : σ →₀ ℕ} (hd : 0 ≺[m] d) : + m.degree f ≺[m] d ↔ ∀ c ∈ f.support, c ≺[m] d := by + simp only [map_zero] at hd + unfold degree + simp only [AddEquiv.apply_symm_apply] + exact Finset.sup_lt_iff hd + +theorem le_degree {f : MvPolynomial σ R} {d : σ →₀ ℕ} (hd : d ∈ f.support) : + d ≼[m] m.degree f := by + unfold degree + simp only [AddEquiv.apply_symm_apply, Finset.le_sup hd] + +theorem coeff_eq_zero_of_lt {f : MvPolynomial σ R} {d : σ →₀ ℕ} (hd : m.degree f ≺[m] d) : + f.coeff d = 0 := by + rw [← not_le] at hd + by_contra hf + apply hd (m.le_degree (mem_support_iff.mpr hf)) +theorem lCoeff_ne_zero_iff {f : MvPolynomial σ R} : + m.lCoeff f ≠ 0 ↔ f ≠ 0 := by + constructor + · rw [not_imp_not] + intro hf + rw [hf, lCoeff_zero] + · intro hf + rw [← support_nonempty] at hf + rw [lCoeff, ← mem_support_iff, degree] + suffices f.support.sup m.toSyn ∈ m.toSyn '' f.support by + obtain ⟨d, hd, hd'⟩ := this + rw [← hd', AddEquiv.symm_apply_apply] + exact hd + exact Finset.sup_mem_of_nonempty hf + +@[simp] +theorem lCoeff_eq_zero_iff {f : MvPolynomial σ R} : + lCoeff m f = 0 ↔ f = 0 := by + simp only [← not_iff_not, lCoeff_ne_zero_iff] + +theorem coeff_degree_ne_zero_iff {f : MvPolynomial σ R} : + f.coeff (m.degree f) ≠ 0 ↔ f ≠ 0 := + m.lCoeff_ne_zero_iff + +@[simp] +theorem coeff_degree_eq_zero_iff {f : MvPolynomial σ R} : + f.coeff (m.degree f) = 0 ↔ f = 0 := + m.lCoeff_eq_zero_iff + +theorem degree_eq_zero_iff_totalDegree_eq_zero {f : MvPolynomial σ R} : + m.degree f = 0 ↔ f.totalDegree = 0 := by + rw [← m.toSyn.injective.eq_iff] + rw [map_zero, ← m.bot_eq_zero, eq_bot_iff, m.bot_eq_zero, ← m.toSyn.map_zero] + rw [degree_le_iff] + rw [totalDegree_eq_zero_iff] + apply forall_congr' + intro d + apply imp_congr (rfl.to_iff) + rw [map_zero, ← m.bot_eq_zero, ← eq_bot_iff, m.bot_eq_zero] + simp only [EmbeddingLike.map_eq_zero_iff] + exact Finsupp.ext_iff + +@[simp] +theorem degree_C (r : R) : + m.degree (C r) = 0 := by + rw [degree_eq_zero_iff_totalDegree_eq_zero, totalDegree_C] + +theorem degree_add_le {f g : MvPolynomial σ R} : + m.toSyn (m.degree (f + g)) ≤ m.toSyn (m.degree f) ⊔ m.toSyn (m.degree g) := by + conv_rhs => rw [← m.toSyn.apply_symm_apply (_ ⊔ _)] + rw [degree_le_iff] + simp only [AddEquiv.apply_symm_apply, le_sup_iff] + intro b hb + by_cases hf : b ∈ f.support + · left + exact m.le_degree hf + · right + apply m.le_degree + simp only [not_mem_support_iff] at hf + simpa only [mem_support_iff, coeff_add, hf, zero_add] using hb + +theorem degree_add_of_lt {f g : MvPolynomial σ R} (h : m.degree g ≺[m] m.degree f) : + m.degree (f + g) = m.degree f := by + apply m.toSyn.injective + apply le_antisymm + · apply le_trans degree_add_le + simp only [sup_le_iff, le_refl, true_and, le_of_lt h] + · apply le_degree + rw [mem_support_iff, coeff_add, m.coeff_eq_zero_of_lt h, add_zero, ← lCoeff, lCoeff_ne_zero_iff] + intro hf + rw [← not_le, hf] at h + apply h + simp only [degree_zero, map_zero] + apply bot_le + +theorem lCoeff_add_of_lt {f g : MvPolynomial σ R} (h : m.degree g ≺[m] m.degree f) : + m.lCoeff (f + g) = m.lCoeff f := by + simp only [lCoeff, m.degree_add_of_lt h, coeff_add, coeff_eq_zero_of_lt h, add_zero] + +theorem degree_add_of_ne {f g : MvPolynomial σ R} + (h : m.degree f ≠ m.degree g) : + m.toSyn (m.degree (f + g)) = m.toSyn (m.degree f) ⊔ m.toSyn (m.degree g) := by + by_cases h' : m.degree g ≺[m] m.degree f + · simp [degree_add_of_lt h', left_eq_sup, le_of_lt h'] + · rw [not_lt, le_iff_eq_or_lt, Classical.or_iff_not_imp_left, EmbeddingLike.apply_eq_iff_eq] at h' + rw [add_comm, degree_add_of_lt (h' h), right_eq_sup] + simp only [le_of_lt (h' h)] + +theorem degree_mul_le {f g : MvPolynomial σ R} : + m.degree (f * g) ≼[m] m.degree f + m.degree g := by + classical + rw [degree_le_iff] + intro c + rw [← not_lt, mem_support_iff, not_imp_not] + intro hc + rw [coeff_mul] + apply Finset.sum_eq_zero + rintro ⟨d, e⟩ hde + simp only [Finset.mem_antidiagonal] at hde + dsimp only + by_cases hd : m.degree f ≺[m] d + · rw [m.coeff_eq_zero_of_lt hd, zero_mul] + · suffices m.degree g ≺[m] e by + rw [m.coeff_eq_zero_of_lt this, mul_zero] + simp only [not_lt] at hd + apply lt_of_add_lt_add_left (a := m.toSyn d) + simp only [← map_add, hde] + apply lt_of_le_of_lt _ hc + simp only [map_add] + exact add_le_add_right hd _ + +/-- Multiplicativity of leading coefficients -/ +theorem coeff_mul_of_degree_add {f g : MvPolynomial σ R} : + (f * g).coeff (m.degree f + m.degree g) = m.lCoeff f * m.lCoeff g := by + classical + rw [coeff_mul] + rw [Finset.sum_eq_single (m.degree f, m.degree g)] + · rfl + · rintro ⟨c, d⟩ hcd h + simp only [Finset.mem_antidiagonal] at hcd + by_cases hf : m.degree f ≺[m] c + · rw [m.coeff_eq_zero_of_lt hf, zero_mul] + · suffices m.degree g ≺[m] d by + rw [coeff_eq_zero_of_lt this, mul_zero] + apply lt_of_add_lt_add_left (a := m.toSyn c) + simp only [← map_add, hcd] + simp only [map_add] + rw [← not_le] + intro h'; apply hf + simp only [le_iff_eq_or_lt] at h' + cases h' with + | inl h' => + simp only [← map_add, EmbeddingLike.apply_eq_iff_eq, add_left_inj] at h' + exfalso + apply h + simp only [h', Prod.mk.injEq, true_and] + simpa [h'] using hcd + | inr h' => + exact lt_of_add_lt_add_right h' + · simp + +/-- Multiplicativity of leading coefficients -/ +theorem degree_mul_of_isRegular_left {f g : MvPolynomial σ R} + (hf : IsRegular (m.lCoeff f)) (hg : g ≠ 0) : + m.degree (f * g) = m.degree f + m.degree g := by + apply m.toSyn.injective + apply le_antisymm degree_mul_le + apply le_degree + rw [mem_support_iff, coeff_mul_of_degree_add] + simp only [ne_eq, hf, IsRegular.left, IsLeftRegular.mul_left_eq_zero_iff, + lCoeff_eq_zero_iff] + exact hg + +/-- Multiplicativity of leading coefficients -/ +theorem lCoeff_mul_of_isRegular_left {f g : MvPolynomial σ R} + (hf : IsRegular (m.lCoeff f)) (hg : g ≠ 0) : + m.lCoeff (f * g) = m.lCoeff f * m.lCoeff g := by + simp only [lCoeff, degree_mul_of_isRegular_left hf hg, coeff_mul_of_degree_add] + +/-- Multiplicativity of leading coefficients -/ +theorem degree_mul_of_isRegular_right {f g : MvPolynomial σ R} + (hf : f ≠ 0) (hg : IsRegular (m.lCoeff g)) : + m.degree (f * g) = m.degree f + m.degree g := by + rw [mul_comm, m.degree_mul_of_isRegular_left hg hf, add_comm] + +/-- Multiplicativity of leading coefficients -/ +theorem lCoeff_mul_of_isRegular_right {f g : MvPolynomial σ R} + (hf : f ≠ 0) (hg : IsRegular (m.lCoeff g)) : + m.lCoeff (f * g) = m.lCoeff f * m.lCoeff g := by + simp only [lCoeff, degree_mul_of_isRegular_right hf hg, coeff_mul_of_degree_add] + +/-- Degree of product -/ +theorem degree_mul [IsDomain R] {f g : MvPolynomial σ R} (hf : f ≠ 0) (hg : g ≠ 0) : + m.degree (f * g) = m.degree f + m.degree g := + degree_mul_of_isRegular_left (isRegular_of_ne_zero (lCoeff_ne_zero_iff.mpr hf)) hg + +/-- Degree of of product -/ +theorem degree_mul_of_nonzero_mul [IsDomain R] {f g : MvPolynomial σ R} (hfg : f * g ≠ 0) : + m.degree (f * g) = m.degree f + m.degree g := + degree_mul (left_ne_zero_of_mul hfg) (right_ne_zero_of_mul hfg) + +/-- Multiplicativity of leading coefficients -/ +theorem lCoeff_mul [IsDomain R] {f g : MvPolynomial σ R} + (hf : f ≠ 0) (hg : g ≠ 0) : + m.lCoeff (f * g) = m.lCoeff f * m.lCoeff g := by + rw [lCoeff, degree_mul hf hg, ← coeff_mul_of_degree_add] + +theorem degree_smul_le {r : R} {f : MvPolynomial σ R} : + m.degree (r • f) ≼[m] m.degree f := by + rw [smul_eq_C_mul] + apply le_of_le_of_eq degree_mul_le + simp + +theorem degree_smul {r : R} (hr : IsRegular r) {f : MvPolynomial σ R} : + m.degree (r • f) = m.degree f := by + by_cases hf : f = 0 + · simp [hf] + apply m.toSyn.injective + apply le_antisymm degree_smul_le + apply le_degree + simp only [mem_support_iff, smul_eq_C_mul] + rw [← zero_add (degree m f), ← degree_C r, coeff_mul_of_degree_add] + simp [lCoeff, hr.left.mul_left_eq_zero_iff, hf] + +theorem eq_C_of_degree_eq_zero {f : MvPolynomial σ R} (hf : m.degree f = 0) : + f = C (m.lCoeff f) := by + ext d + simp only [lCoeff, hf] + classical + by_cases hd : d = 0 + · simp [hd] + · rw [coeff_C, if_neg (Ne.symm hd)] + apply coeff_eq_zero_of_lt (m := m) + rw [hf, map_zero, lt_iff_le_and_ne, ne_eq, eq_comm, EmbeddingLike.map_eq_zero_iff] + exact ⟨bot_le, hd⟩ + +end Semiring + +section Ring + +variable {R : Type*} [CommRing R] + +@[simp] +theorem degree_neg {f : MvPolynomial σ R} : + m.degree (-f) = m.degree f := by + unfold degree + rw [support_neg] + +theorem degree_sub_le {f g : MvPolynomial σ R} : + m.toSyn (m.degree (f - g)) ≤ m.toSyn (m.degree f) ⊔ m.toSyn (m.degree g) := by + rw [sub_eq_add_neg] + apply le_of_le_of_eq m.degree_add_le + rw [degree_neg] + +theorem degree_sub_of_lt {f g : MvPolynomial σ R} (h : m.degree g ≺[m] m.degree f) : + m.degree (f - g) = m.degree f := by + rw [sub_eq_add_neg] + apply degree_add_of_lt + simp only [degree_neg, h] + +theorem lCoeff_sub_of_lt {f g : MvPolynomial σ R} (h : m.degree g ≺[m] m.degree f) : + m.lCoeff (f - g) = m.lCoeff f := by + rw [sub_eq_add_neg] + apply lCoeff_add_of_lt + simp only [degree_neg, h] + +end Ring + +section Field + +variable {R : Type*} [Field R] + +theorem lCoeff_is_unit_iff {f : MvPolynomial σ R} : + IsUnit (m.lCoeff f) ↔ f ≠ 0 := by + simp only [isUnit_iff_ne_zero, ne_eq, lCoeff_eq_zero_iff] + +end Field + +end MonomialOrder From da0cfc4d741436ce977bb835884026094ff2a16f Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Wed, 4 Dec 2024 09:24:04 +0000 Subject: [PATCH 505/829] perf: speedup using `Submodule.restrictScalars_mem` instead of relying on defeq (#19711) --- Mathlib/RingTheory/Ideal/Cotangent.lean | 1 + Mathlib/RingTheory/Kaehler/Basic.lean | 2 +- 2 files changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/Ideal/Cotangent.lean b/Mathlib/RingTheory/Ideal/Cotangent.lean index 04fd59250aa5c..9f2e89438bbbc 100644 --- a/Mathlib/RingTheory/Ideal/Cotangent.lean +++ b/Mathlib/RingTheory/Ideal/Cotangent.lean @@ -196,6 +196,7 @@ def mapCotangent (I₁ : Ideal A) (I₂ : Ideal B) (f : A →ₐ[R] B) (h : I₁ ((I₂ • ⊤ : Submodule B I₂).restrictScalars R) ?_ ?_ · exact f.toLinearMap.restrict (p := I₁.restrictScalars R) (q := I₂.restrictScalars R) h · intro x hx + rw [Submodule.restrictScalars_mem] at hx refine Submodule.smul_induction_on hx ?_ (fun _ _ ↦ add_mem) rintro a ha ⟨b, hb⟩ - simp only [SetLike.mk_smul_mk, smul_eq_mul, Submodule.mem_comap, Submodule.restrictScalars_mem] diff --git a/Mathlib/RingTheory/Kaehler/Basic.lean b/Mathlib/RingTheory/Kaehler/Basic.lean index 96de1da9e85a4..4fa13a26c251e 100644 --- a/Mathlib/RingTheory/Kaehler/Basic.lean +++ b/Mathlib/RingTheory/Kaehler/Basic.lean @@ -259,7 +259,7 @@ def Derivation.liftKaehlerDifferential (D : Derivation R S M) : Ω[S⁄R] →ₗ · exact D.tensorProductTo.comp ((KaehlerDifferential.ideal R S).subtype.restrictScalars S) · intro x hx rw [LinearMap.mem_ker] - refine Submodule.smul_induction_on hx ?_ ?_ + refine Submodule.smul_induction_on ((Submodule.restrictScalars_mem _ _ _).mp hx) ?_ ?_ · rintro x hx y - rw [RingHom.mem_ker] at hx dsimp From d7abea93d01641d025a6fdcaa8bc0298f17cd9ee Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Wed, 4 Dec 2024 10:22:02 +0000 Subject: [PATCH 506/829] chore: golf IsCycles.other_adj_of_adj (#19710) --- Mathlib/Combinatorics/SimpleGraph/Matching.lean | 8 +++----- 1 file changed, 3 insertions(+), 5 deletions(-) diff --git a/Mathlib/Combinatorics/SimpleGraph/Matching.lean b/Mathlib/Combinatorics/SimpleGraph/Matching.lean index 5d6ee34bb110b..1e27af2e43ad5 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Matching.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Matching.lean @@ -333,11 +333,9 @@ to the same vertex. lemma IsCycles.other_adj_of_adj (h : G.IsCycles) (hadj : G.Adj v w) : ∃ w', w ≠ w' ∧ G.Adj v w' := by simp_rw [← SimpleGraph.mem_neighborSet] at hadj ⊢ - cases' (G.neighborSet v).eq_empty_or_nonempty with hl hr - · exact ((Set.eq_empty_iff_forall_not_mem.mp hl) _ hadj).elim - · have := h hr - obtain ⟨w', hww'⟩ := (G.neighborSet v).exists_ne_of_one_lt_ncard (by omega) w - exact ⟨w', ⟨hww'.2.symm, hww'.1⟩⟩ + have := h ⟨w, hadj⟩ + obtain ⟨w', hww'⟩ := (G.neighborSet v).exists_ne_of_one_lt_ncard (by omega) w + exact ⟨w', ⟨hww'.2.symm, hww'.1⟩⟩ open scoped symmDiff From bbaf1a0ef4ec577ac6ee418922b59690c66ba196 Mon Sep 17 00:00:00 2001 From: Christopher Hoskin Date: Wed, 4 Dec 2024 11:42:34 +0000 Subject: [PATCH 507/829] feat(Logic/Basic): xor_iff_or_and_not_and (#19703) Characterises `Xor'` as "one or the other but not both". c.f. `symmDiff_eq'` in `Order/SymmDiff` Co-authored-by: Christopher Hoskin --- Mathlib/Logic/Basic.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index 13c8fd29ba91a..d63de1913084e 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -353,6 +353,10 @@ theorem xor_iff_iff_not : Xor' a b ↔ (a ↔ ¬b) := by simp only [← @xor_not theorem xor_iff_not_iff' : Xor' a b ↔ (¬a ↔ b) := by simp only [← @xor_not_left _ b, not_not] +theorem xor_iff_or_and_not_and (a b : Prop) : Xor' a b ↔ (a ∨ b) ∧ (¬ (a ∧ b)) := by + rw [Xor', or_and_right, not_and_or, and_or_left, and_not_self_iff, false_or, + and_or_left, and_not_self_iff, or_false] + end Propositional /-! ### Declarations about equality -/ From a38db992d088e602ba99cdf94f77435bf23e9b8e Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Wed, 4 Dec 2024 12:33:11 +0000 Subject: [PATCH 508/829] refactor: generalize `SMul (Ideal R) (Submodule R M)` to `SMul (Submodule R A) (Submodule R M)` (#18419) and redefine `Mul (Submodule R A) (Submodule R A)` using the latter. --- Mathlib/Algebra/Algebra/Operations.lean | 168 ++++++++++++++++++----- Mathlib/RingTheory/Ideal/Operations.lean | 92 +------------ 2 files changed, 138 insertions(+), 122 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Operations.lean b/Mathlib/Algebra/Algebra/Operations.lean index 7c503761b174e..24dddc1deceec 100644 --- a/Mathlib/Algebra/Algebra/Operations.lean +++ b/Mathlib/Algebra/Algebra/Operations.lean @@ -93,24 +93,138 @@ theorem one_le {P : Submodule R A} : (1 : Submodule R A) ≤ P ↔ (1 : A) ∈ P -- Porting note: simpa no longer closes refl goals, so added `SetLike.mem_coe` simp only [one_eq_span, span_le, Set.singleton_subset_iff, SetLike.mem_coe] +variable {M : Type*} [AddCommMonoid M] [Module R M] [Module A M] [IsScalarTower R A M] + +instance : SMul (Submodule R A) (Submodule R M) where + smul A' M' := + { __ := A'.toAddSubmonoid • M'.toAddSubmonoid + smul_mem' := fun r m hm ↦ AddSubmonoid.smul_induction_on hm + (fun a ha m hm ↦ by rw [← smul_assoc]; exact AddSubmonoid.smul_mem_smul (A'.smul_mem r ha) hm) + fun m₁ m₂ h₁ h₂ ↦ by rw [smul_add]; exact (A'.1 • M'.1).add_mem h₁ h₂ } + +section + +variable {I J : Submodule R A} {N P : Submodule R M} + +theorem smul_toAddSubmonoid : (I • N).toAddSubmonoid = I.toAddSubmonoid • N.toAddSubmonoid := rfl + +theorem smul_mem_smul {r} {n} (hr : r ∈ I) (hn : n ∈ N) : r • n ∈ I • N := + AddSubmonoid.smul_mem_smul hr hn + +theorem smul_le : I • N ≤ P ↔ ∀ r ∈ I, ∀ n ∈ N, r • n ∈ P := + AddSubmonoid.smul_le + +@[simp, norm_cast] +lemma coe_set_smul : (I : Set A) • N = I • N := + set_smul_eq_of_le _ _ _ + (fun _ _ hr hx ↦ smul_mem_smul hr hx) + (smul_le.mpr fun _ hr _ hx ↦ mem_set_smul_of_mem_mem hr hx) + +@[elab_as_elim] +theorem smul_induction_on {p : M → Prop} {x} (H : x ∈ I • N) (smul : ∀ r ∈ I, ∀ n ∈ N, p (r • n)) + (add : ∀ x y, p x → p y → p (x + y)) : p x := + AddSubmonoid.smul_induction_on H smul add + +/-- Dependent version of `Submodule.smul_induction_on`. -/ +@[elab_as_elim] +theorem smul_induction_on' {x : M} (hx : x ∈ I • N) {p : ∀ x, x ∈ I • N → Prop} + (smul : ∀ (r : A) (hr : r ∈ I) (n : M) (hn : n ∈ N), p (r • n) (smul_mem_smul hr hn)) + (add : ∀ x hx y hy, p x hx → p y hy → p (x + y) (add_mem ‹_› ‹_›)) : p x hx := by + refine Exists.elim ?_ fun (h : x ∈ I • N) (H : p x h) ↦ H + exact smul_induction_on hx (fun a ha x hx ↦ ⟨_, smul _ ha _ hx⟩) + fun x y ⟨_, hx⟩ ⟨_, hy⟩ ↦ ⟨_, add _ _ _ _ hx hy⟩ + +theorem smul_mono (hij : I ≤ J) (hnp : N ≤ P) : I • N ≤ J • P := + AddSubmonoid.smul_le_smul hij hnp + +theorem smul_mono_left (h : I ≤ J) : I • N ≤ J • N := + smul_mono h le_rfl + +instance : CovariantClass (Submodule R A) (Submodule R M) HSMul.hSMul LE.le := + ⟨fun _ _ => smul_mono le_rfl⟩ + +@[deprecated smul_mono_right (since := "2024-03-31")] +protected theorem smul_mono_right (h : N ≤ P) : I • N ≤ I • P := + _root_.smul_mono_right I h + +variable (I J N P) + +@[simp] +theorem smul_bot : I • (⊥ : Submodule R M) = ⊥ := + toAddSubmonoid_injective <| AddSubmonoid.addSubmonoid_smul_bot _ + +@[simp] +theorem bot_smul : (⊥ : Submodule R A) • N = ⊥ := + le_bot_iff.mp <| smul_le.mpr <| by rintro _ rfl _ _; rw [zero_smul]; exact zero_mem _ + +theorem smul_sup : I • (N ⊔ P) = I • N ⊔ I • P := + toAddSubmonoid_injective <| by + simp only [smul_toAddSubmonoid, sup_toAddSubmonoid, AddSubmonoid.addSubmonoid_smul_sup] + +theorem sup_smul : (I ⊔ J) • N = I • N ⊔ J • N := + le_antisymm (smul_le.mpr fun mn hmn p hp ↦ by + obtain ⟨m, hm, n, hn, rfl⟩ := mem_sup.mp hmn + rw [add_smul]; exact add_mem_sup (smul_mem_smul hm hp) <| smul_mem_smul hn hp) + (sup_le (smul_mono_left le_sup_left) <| smul_mono_left le_sup_right) + +protected theorem smul_assoc {B} [Semiring B] [Module R B] [Module A B] [Module B M] + [IsScalarTower R A B] [IsScalarTower R B M] [IsScalarTower A B M] + (I : Submodule R A) (J : Submodule R B) (N : Submodule R M) : + (I • J) • N = I • J • N := + le_antisymm + (smul_le.2 fun _ hrsij t htn ↦ smul_induction_on hrsij + (fun r hr s hs ↦ smul_assoc r s t ▸ smul_mem_smul hr (smul_mem_smul hs htn)) + fun x y ↦ (add_smul x y t).symm ▸ add_mem) + (smul_le.2 fun r hr _ hsn ↦ smul_induction_on hsn + (fun j hj n hn ↦ (smul_assoc r j n).symm ▸ smul_mem_smul (smul_mem_smul hr hj) hn) + fun m₁ m₂ ↦ (smul_add r m₁ m₂) ▸ add_mem) + +@[deprecated smul_inf_le (since := "2024-03-31")] +protected theorem smul_inf_le (M₁ M₂ : Submodule R M) : + I • (M₁ ⊓ M₂) ≤ I • M₁ ⊓ I • M₂ := smul_inf_le _ _ _ + +theorem smul_iSup {ι : Sort*} {I : Submodule R A} {t : ι → Submodule R M} : + I • (⨆ i, t i)= ⨆ i, I • t i := + toAddSubmonoid_injective <| by + simp only [smul_toAddSubmonoid, iSup_toAddSubmonoid, AddSubmonoid.smul_iSup] + +theorem iSup_smul {ι : Sort*} {t : ι → Submodule R A} {N : Submodule R M} : + (⨆ i, t i) • N = ⨆ i, t i • N := + le_antisymm (smul_le.mpr fun t ht s hs ↦ iSup_induction _ (C := (· • s ∈ _)) ht + (fun i t ht ↦ mem_iSup_of_mem i <| smul_mem_smul ht hs) + (by simp_rw [zero_smul]; apply zero_mem) fun x y ↦ by simp_rw [add_smul]; apply add_mem) + (iSup_le fun i ↦ Submodule.smul_mono_left <| le_iSup _ i) + +@[deprecated smul_iInf_le (since := "2024-03-31")] +protected theorem smul_iInf_le {ι : Sort*} {I : Submodule R A} {t : ι → Submodule R M} : + I • iInf t ≤ ⨅ i, I • t i := + smul_iInf_le + +protected theorem one_smul : (1 : Submodule R A) • N = N := by + refine le_antisymm (smul_le.mpr fun r hr m hm ↦ ?_) fun m hm ↦ ?_ + · obtain ⟨r, rfl⟩ := hr + rw [LinearMap.toSpanSingleton_apply, smul_one_smul]; exact N.smul_mem r hm + · rw [← one_smul A m]; exact smul_mem_smul (one_le.mp le_rfl) hm + +theorem smul_subset_smul : (↑I : Set A) • (↑N : Set M) ⊆ (↑(I • N) : Set M) := + AddSubmonoid.smul_subset_smul + +end + variable [IsScalarTower R A A] /-- Multiplication of sub-R-modules of an R-module A that is also a semiring. The submodule `M * N` consists of finite sums of elements `m * n` for `m ∈ M` and `n ∈ N`. -/ instance mul : Mul (Submodule R A) where - mul M N := - { __ := M.toAddSubmonoid * N.toAddSubmonoid - smul_mem' := fun r a ha ↦ AddSubmonoid.mul_induction_on ha - (fun m h n h' ↦ by rw [← smul_mul_assoc]; exact AddSubmonoid.mul_mem_mul (M.smul_mem r h) h') - fun a₁ a₂ h₁ h₂ ↦ by rw [smul_add]; apply (M.1 * N.1).add_mem h₁ h₂ } + mul := (· • ·) variable (S T : Set A) {M N P Q : Submodule R A} {m n : A} theorem mul_mem_mul (hm : m ∈ M) (hn : n ∈ N) : m * n ∈ M * N := - AddSubmonoid.mul_mem_mul hm hn + smul_mem_smul hm hn theorem mul_le : M * N ≤ P ↔ ∀ m ∈ M, ∀ n ∈ N, m * n ∈ P := - AddSubmonoid.mul_le + smul_le theorem mul_toAddSubmonoid (M N : Submodule R A) : (M * N).toAddSubmonoid = M.toAddSubmonoid * N.toAddSubmonoid := rfl @@ -118,46 +232,40 @@ theorem mul_toAddSubmonoid (M N : Submodule R A) : @[elab_as_elim] protected theorem mul_induction_on {C : A → Prop} {r : A} (hr : r ∈ M * N) (hm : ∀ m ∈ M, ∀ n ∈ N, C (m * n)) (ha : ∀ x y, C x → C y → C (x + y)) : C r := - AddSubmonoid.mul_induction_on hr hm ha + smul_induction_on hr hm ha /-- A dependent version of `mul_induction_on`. -/ @[elab_as_elim] protected theorem mul_induction_on' {C : ∀ r, r ∈ M * N → Prop} (mem_mul_mem : ∀ m (hm : m ∈ M) n (hn : n ∈ N), C (m * n) (mul_mem_mul hm hn)) (add : ∀ x hx y hy, C x hx → C y hy → C (x + y) (add_mem hx hy)) {r : A} (hr : r ∈ M * N) : - C r hr := by - refine Exists.elim ?_ fun (hr : r ∈ M * N) (hc : C r hr) ↦ hc - exact Submodule.mul_induction_on hr - (fun x hx y hy ↦ ⟨_, mem_mul_mem _ hx _ hy⟩) - fun x y ⟨_, hx⟩ ⟨_, hy⟩ ↦ ⟨_, add _ _ _ _ hx hy⟩ + C r hr := + smul_induction_on' hr mem_mul_mem add variable (M) @[simp] theorem mul_bot : M * ⊥ = ⊥ := - toAddSubmonoid_injective (AddSubmonoid.mul_bot _) + smul_bot _ @[simp] theorem bot_mul : ⊥ * M = ⊥ := - toAddSubmonoid_injective (AddSubmonoid.bot_mul _) + bot_smul _ -protected theorem one_mul : (1 : Submodule R A) * M = M := by - refine le_antisymm (mul_le.mpr fun r hr m hm ↦ ?_) fun m hm ↦ ?_ - · obtain ⟨r, rfl⟩ := hr - rw [LinearMap.toSpanSingleton_apply, smul_one_mul]; exact M.smul_mem r hm - · rw [← one_mul m]; exact mul_mem_mul (one_le.mp le_rfl) hm +protected theorem one_mul : (1 : Submodule R A) * M = M := + Submodule.one_smul _ variable {M} @[mono] theorem mul_le_mul (hmp : M ≤ P) (hnq : N ≤ Q) : M * N ≤ P * Q := - AddSubmonoid.mul_le_mul hmp hnq + smul_mono hmp hnq theorem mul_le_mul_left (h : M ≤ N) : M * P ≤ N * P := - AddSubmonoid.mul_le_mul_left h + smul_mono_left h theorem mul_le_mul_right (h : N ≤ P) : M * N ≤ M * P := - AddSubmonoid.mul_le_mul_right h + smul_mono_right _ h theorem mul_comm_of_commute (h : ∀ m ∈ M, ∀ n ∈ N, Commute m n) : M * N = N * M := toAddSubmonoid_injective <| AddSubmonoid.mul_comm_of_commute h @@ -165,15 +273,13 @@ theorem mul_comm_of_commute (h : ∀ m ∈ M, ∀ n ∈ N, Commute m n) : M * N variable (M N P) theorem mul_sup : M * (N ⊔ P) = M * N ⊔ M * P := - toAddSubmonoid_injective <| by - simp only [mul_toAddSubmonoid, sup_toAddSubmonoid, AddSubmonoid.mul_sup] + smul_sup _ _ _ theorem sup_mul : (M ⊔ N) * P = M * P ⊔ N * P := - toAddSubmonoid_injective <| by - simp only [mul_toAddSubmonoid, sup_toAddSubmonoid, AddSubmonoid.sup_mul] + sup_smul _ _ _ theorem mul_subset_mul : (↑M : Set A) * (↑N : Set A) ⊆ (↑(M * N) : Set A) := - AddSubmonoid.mul_subset_mul + smul_subset_smul _ _ lemma restrictScalars_mul {A B C} [Semiring A] [Semiring B] [Semiring C] [SMul A B] [Module A C] [Module B C] [IsScalarTower A C C] [IsScalarTower B C C] @@ -184,12 +290,10 @@ lemma restrictScalars_mul {A B C} [Semiring A] [Semiring B] [Semiring C] variable {ι : Sort uι} theorem iSup_mul (s : ι → Submodule R A) (t : Submodule R A) : (⨆ i, s i) * t = ⨆ i, s i * t := - toAddSubmonoid_injective <| by - simp only [mul_toAddSubmonoid, iSup_toAddSubmonoid, AddSubmonoid.iSup_mul] + iSup_smul theorem mul_iSup (t : Submodule R A) (s : ι → Submodule R A) : (t * ⨆ i, s i) = ⨆ i, t * s i := - toAddSubmonoid_injective <| by - simp only [mul_toAddSubmonoid, iSup_toAddSubmonoid, AddSubmonoid.mul_iSup] + smul_iSup /-- Sub-`R`-modules of an `R`-module form an idempotent semiring. -/ instance : NonUnitalSemiring (Submodule R A) where diff --git a/Mathlib/RingTheory/Ideal/Operations.lean b/Mathlib/RingTheory/Ideal/Operations.lean index d627344ccb869..cb663d99cc855 100644 --- a/Mathlib/RingTheory/Ideal/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Operations.lean @@ -56,116 +56,28 @@ section Semiring variable [Semiring R] [AddCommMonoid M] [Module R M] -instance : SMul (Ideal R) (Submodule R M) where - smul I N := - { __ := I.toAddSubmonoid • N.toAddSubmonoid - smul_mem' := fun r m hm ↦ AddSubmonoid.smul_induction_on hm - (fun m hm n ↦ by rw [smul_smul]; exact AddSubmonoid.smul_mem_smul <| I.smul_mem _ hm) - fun m₁ m₂ h₁ h₂ ↦ by rw [smul_add]; exact (I.1 • N.1).add_mem h₁ h₂ } - /-- This duplicates the global `smul_eq_mul`, but doesn't have to unfold anywhere near as much to apply. -/ protected theorem _root_.Ideal.smul_eq_mul (I J : Ideal R) : I • J = I * J := rfl -variable {I J : Ideal R} {N P : Submodule R M} - -theorem smul_toAddSubmonoid : (I • N).toAddSubmonoid = I.toAddSubmonoid • N.toAddSubmonoid := rfl - -theorem smul_mem_smul {r} {n} (hr : r ∈ I) (hn : n ∈ N) : r • n ∈ I • N := - AddSubmonoid.smul_mem_smul hr hn - -theorem smul_le : I • N ≤ P ↔ ∀ r ∈ I, ∀ n ∈ N, r • n ∈ P := - AddSubmonoid.smul_le - -@[simp, norm_cast] -lemma coe_set_smul : (I : Set R) • N = I • N := - set_smul_eq_of_le _ _ _ - (fun _ _ hr hx ↦ smul_mem_smul hr hx) - (smul_le.mpr fun _ hr _ hx ↦ mem_set_smul_of_mem_mem hr hx) - -@[elab_as_elim] -theorem smul_induction_on {p : M → Prop} {x} (H : x ∈ I • N) (smul : ∀ r ∈ I, ∀ n ∈ N, p (r • n)) - (add : ∀ x y, p x → p y → p (x + y)) : p x := - AddSubmonoid.smul_induction_on H smul add - -/-- Dependent version of `Submodule.smul_induction_on`. -/ -@[elab_as_elim] -theorem smul_induction_on' {x : M} (hx : x ∈ I • N) {p : ∀ x, x ∈ I • N → Prop} - (smul : ∀ (r : R) (hr : r ∈ I) (n : M) (hn : n ∈ N), p (r • n) (smul_mem_smul hr hn)) - (add : ∀ x hx y hy, p x hx → p y hy → p (x + y) (add_mem ‹_› ‹_›)) : p x hx := by - refine Exists.elim ?_ fun (h : x ∈ I • N) (H : p x h) ↦ H - exact smul_induction_on hx (fun a ha x hx ↦ ⟨_, smul _ ha _ hx⟩) - fun x y ⟨_, hx⟩ ⟨_, hy⟩ ↦ ⟨_, add _ _ _ _ hx hy⟩ +variable {I : Ideal R} {N : Submodule R M} theorem smul_le_right : I • N ≤ N := smul_le.2 fun r _ _ ↦ N.smul_mem r -theorem smul_mono (hij : I ≤ J) (hnp : N ≤ P) : I • N ≤ J • P := - AddSubmonoid.smul_le_smul hij hnp - -theorem smul_mono_left (h : I ≤ J) : I • N ≤ J • N := - smul_mono h le_rfl - -instance : CovariantClass (Ideal R) (Submodule R M) HSMul.hSMul LE.le := - ⟨fun _ _ => smul_mono le_rfl⟩ - -@[deprecated smul_mono_right (since := "2024-03-31")] -protected theorem smul_mono_right (h : N ≤ P) : I • N ≤ I • P := - _root_.smul_mono_right I h - theorem map_le_smul_top (I : Ideal R) (f : R →ₗ[R] M) : Submodule.map f I ≤ I • (⊤ : Submodule R M) := by rintro _ ⟨y, hy, rfl⟩ rw [← mul_one y, ← smul_eq_mul, f.map_smul] exact smul_mem_smul hy mem_top -variable (I J N P) - -@[simp] -theorem smul_bot : I • (⊥ : Submodule R M) = ⊥ := - toAddSubmonoid_injective <| AddSubmonoid.addSubmonoid_smul_bot _ - -@[simp] -theorem bot_smul : (⊥ : Ideal R) • N = ⊥ := - le_bot_iff.mp <| smul_le.mpr <| by rintro _ rfl _ _; rw [zero_smul]; exact zero_mem _ +variable (I N) @[simp] theorem top_smul : (⊤ : Ideal R) • N = N := le_antisymm smul_le_right fun r hri => one_smul R r ▸ smul_mem_smul mem_top hri -theorem smul_sup : I • (N ⊔ P) = I • N ⊔ I • P := - toAddSubmonoid_injective <| by - simp only [smul_toAddSubmonoid, sup_toAddSubmonoid, AddSubmonoid.addSubmonoid_smul_sup] - -theorem sup_smul : (I ⊔ J) • N = I • N ⊔ J • N := - le_antisymm (smul_le.mpr fun mn hmn p hp ↦ by - obtain ⟨m, hm, n, hn, rfl⟩ := mem_sup.mp hmn - rw [add_smul]; exact add_mem_sup (smul_mem_smul hm hp) <| smul_mem_smul hn hp) - (sup_le (smul_mono_left le_sup_left) <| smul_mono_left le_sup_right) - -protected theorem smul_assoc : (I • J) • N = I • J • N := - le_antisymm - (smul_le.2 fun _ hrsij t htn ↦ smul_induction_on hrsij - (fun r hr s hs ↦ smul_assoc r s t ▸ smul_mem_smul hr (smul_mem_smul hs htn)) - fun x y ↦ (add_smul x y t).symm ▸ add_mem) - (smul_le.2 fun r hr _ hsn ↦ smul_induction_on hsn - (fun j hj n hn ↦ (smul_assoc r j n).symm ▸ smul_mem_smul (smul_mem_smul hr hj) hn) - fun m₁ m₂ ↦ (smul_add r m₁ m₂) ▸ add_mem) - -@[deprecated smul_inf_le (since := "2024-03-31")] -protected theorem smul_inf_le (M₁ M₂ : Submodule R M) : - I • (M₁ ⊓ M₂) ≤ I • M₁ ⊓ I • M₂ := smul_inf_le _ _ _ - -theorem smul_iSup {ι : Sort*} {I : Ideal R} {t : ι → Submodule R M} : I • iSup t = ⨆ i, I • t i := - toAddSubmonoid_injective <| by - simp only [smul_toAddSubmonoid, iSup_toAddSubmonoid, AddSubmonoid.smul_iSup] - -@[deprecated smul_iInf_le (since := "2024-03-31")] -protected theorem smul_iInf_le {ι : Sort*} {I : Ideal R} {t : ι → Submodule R M} : - I • iInf t ≤ ⨅ i, I • t i := - smul_iInf_le - theorem mem_of_span_top_of_smul_mem (M' : Submodule R M) (s : Set R) (hs : Ideal.span s = ⊤) (x : M) (H : ∀ r : s, (r : R) • x ∈ M') : x ∈ M' := by suffices LinearMap.range (LinearMap.toSpanSingleton R M x) ≤ M' by From 41cbc10ab199e188e201edc5cc23a371a835dad8 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Wed, 4 Dec 2024 17:24:56 +0000 Subject: [PATCH 509/829] feat: unique differentiability of preimages under projection doesn't need the bundle to be smooth (#19682) --- .../Geometry/Manifold/MFDeriv/Tangent.lean | 2 +- .../Manifold/MFDeriv/UniqueDifferential.lean | 86 ++++++++++++++----- 2 files changed, 65 insertions(+), 23 deletions(-) diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Tangent.lean b/Mathlib/Geometry/Manifold/MFDeriv/Tangent.lean index 7afe6dc78bd73..e3df71c1fd846 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Tangent.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Tangent.lean @@ -62,7 +62,7 @@ lemma mfderiv_chartAt_eq_tangentCoordChange {x y : M} (hsrc : x ∈ (chartAt H y the basis also has unique differential. -/ theorem UniqueMDiffOn.tangentBundle_proj_preimage {s : Set M} (hs : UniqueMDiffOn I s) : UniqueMDiffOn I.tangent (π E (TangentSpace I) ⁻¹' s) := - hs.smooth_bundle_preimage _ + hs.bundle_preimage _ /-- To write a linear map between tangent spaces in coordinates amounts to precomposing and postcomposing it with derivatives of extended charts. diff --git a/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean b/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean index e090b101ab9e7..187e4cca99ec5 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean @@ -115,36 +115,78 @@ end open Bundle variable {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] {Z : M → Type*} - [TopologicalSpace (TotalSpace F Z)] [∀ b, TopologicalSpace (Z b)] [∀ b, AddCommMonoid (Z b)] - [∀ b, Module 𝕜 (Z b)] [FiberBundle F Z] [VectorBundle 𝕜 F Z] [SmoothVectorBundle F Z I] + [TopologicalSpace (TotalSpace F Z)] [∀ b, TopologicalSpace (Z b)] [FiberBundle F Z] -theorem Trivialization.mdifferentiable (e : Trivialization F (π F Z)) [MemTrivializationAtlas e] : - e.toPartialHomeomorph.MDifferentiable (I.prod 𝓘(𝕜, F)) (I.prod 𝓘(𝕜, F)) := - ⟨e.contMDiffOn.mdifferentiableOn le_top, e.contMDiffOn_symm.mdifferentiableOn le_top⟩ - -theorem UniqueMDiffWithinAt.smooth_bundle_preimage {p : TotalSpace F Z} +private lemma UniqueMDiffWithinAt.bundle_preimage_aux {p : TotalSpace F Z} + (hs : UniqueMDiffWithinAt I s p.proj) (h's : s ⊆ (trivializationAt F Z p.proj).baseSet) : + UniqueMDiffWithinAt (I.prod 𝓘(𝕜, F)) (π F Z ⁻¹' s) p := by + suffices ((extChartAt I p.proj).symm ⁻¹' s ∩ range I) ×ˢ univ ⊆ + (extChartAt (I.prod 𝓘(𝕜, F)) p).symm ⁻¹' (TotalSpace.proj ⁻¹' s) ∩ range (I.prod 𝓘(𝕜, F)) by + let w := (extChartAt (I.prod 𝓘(𝕜, F)) p p).2 + have A : extChartAt (I.prod 𝓘(𝕜, F)) p p = (extChartAt I p.1 p.1, w) := by + ext + · simp [FiberBundle.chartedSpace_chartAt] + · rfl + simp only [UniqueMDiffWithinAt, A] at hs ⊢ + exact (hs.prod (uniqueDiffWithinAt_univ (x := w))).mono this + rcases p with ⟨x, v⟩ + dsimp + rintro ⟨z, w⟩ ⟨hz, -⟩ + simp only [ModelWithCorners.target_eq, mem_inter_iff, mem_preimage, Function.comp_apply, + mem_range] at hz + simp only [FiberBundle.chartedSpace_chartAt, PartialHomeomorph.coe_trans_symm, mem_inter_iff, + mem_preimage, Function.comp_apply, mem_range] + constructor + · rw [PartialEquiv.prod_symm, PartialEquiv.refl_symm, PartialEquiv.prod_coe, + ModelWithCorners.toPartialEquiv_coe_symm, PartialEquiv.refl_coe, + PartialHomeomorph.prod_symm, PartialHomeomorph.refl_symm, PartialHomeomorph.prod_apply, + PartialHomeomorph.refl_apply] + convert hz.1 + apply Trivialization.proj_symm_apply' + exact h's hz.1 + · rcases hz.2 with ⟨u, rfl⟩ + exact ⟨(u, w), rfl⟩ + +/-- In a fiber bundle, the preimage under the projection of a set with unique differentials +in the base has unique differentials in the bundle. -/ +theorem UniqueMDiffWithinAt.bundle_preimage {p : TotalSpace F Z} (hs : UniqueMDiffWithinAt I s p.proj) : UniqueMDiffWithinAt (I.prod 𝓘(𝕜, F)) (π F Z ⁻¹' s) p := by - set e := trivializationAt F Z p.proj - have hp : p ∈ e.source := FiberBundle.mem_trivializationAt_proj_source - have : UniqueMDiffWithinAt (I.prod 𝓘(𝕜, F)) (s ×ˢ univ) (e p) := by - rw [← Prod.mk.eta (p := e p), FiberBundle.trivializationAt_proj_fst] - exact hs.prod (uniqueMDiffWithinAt_univ _) - rw [← e.left_inv hp] - refine (this.preimage_partialHomeomorph e.mdifferentiable.symm (e.map_source hp)).mono ?_ - rintro y ⟨hy, hys, -⟩ - rwa [PartialHomeomorph.symm_symm, e.coe_coe, e.coe_fst hy] at hys + suffices UniqueMDiffWithinAt (I.prod 𝓘(𝕜, F)) + (π F Z ⁻¹' (s ∩ (trivializationAt F Z p.proj).baseSet)) p from this.mono (by simp) + apply UniqueMDiffWithinAt.bundle_preimage_aux (hs.inter _) inter_subset_right + exact IsOpen.mem_nhds (trivializationAt F Z p.proj).open_baseSet + (FiberBundle.mem_baseSet_trivializationAt' p.proj) + +@[deprecated (since := "2024-12-02")] +alias UniqueMDiffWithinAt.smooth_bundle_preimage := UniqueMDiffWithinAt.bundle_preimage variable (Z) -theorem UniqueMDiffWithinAt.smooth_bundle_preimage' {b : M} (hs : UniqueMDiffWithinAt I s b) +/-- In a fiber bundle, the preimage under the projection of a set with unique differentials +in the base has unique differentials in the bundle. Version with a point `⟨b, x⟩`. -/ +theorem UniqueMDiffWithinAt.bundle_preimage' {b : M} (hs : UniqueMDiffWithinAt I s b) (x : Z b) : UniqueMDiffWithinAt (I.prod 𝓘(𝕜, F)) (π F Z ⁻¹' s) ⟨b, x⟩ := - hs.smooth_bundle_preimage (p := ⟨b, x⟩) + hs.bundle_preimage (p := ⟨b, x⟩) + +@[deprecated (since := "2024-12-02")] +alias UniqueMDiffWithinAt.smooth_bundle_preimage' := UniqueMDiffWithinAt.bundle_preimage' -/-- In a smooth fiber bundle, the preimage under the projection of a set with -unique differential in the basis also has unique differential. -/ -theorem UniqueMDiffOn.smooth_bundle_preimage (hs : UniqueMDiffOn I s) : +/-- In a fiber bundle, the preimage under the projection of a set with unique differentials +in the base has unique differentials in the bundle. -/ +theorem UniqueMDiffOn.bundle_preimage (hs : UniqueMDiffOn I s) : UniqueMDiffOn (I.prod 𝓘(𝕜, F)) (π F Z ⁻¹' s) := fun _p hp ↦ - (hs _ hp).smooth_bundle_preimage + (hs _ hp).bundle_preimage + +@[deprecated (since := "2024-12-02")] +alias UniqueMDiffOn.smooth_bundle_preimage := UniqueMDiffOn.bundle_preimage + +/- TODO: move me to `Mathlib.Geometry.Manifold.VectorBundle.MDifferentiable` once #19636 is in. -/ +variable [∀ b, AddCommMonoid (Z b)] [∀ b, Module 𝕜 (Z b)] [VectorBundle 𝕜 F Z] + +theorem Trivialization.mdifferentiable [SmoothVectorBundle F Z I] + (e : Trivialization F (π F Z)) [MemTrivializationAtlas e] : + e.MDifferentiable (I.prod 𝓘(𝕜, F)) (I.prod 𝓘(𝕜, F)) := + ⟨e.contMDiffOn.mdifferentiableOn le_top, e.contMDiffOn_symm.mdifferentiableOn le_top⟩ end UniqueMDiff From 0e2973b0fb7d17e76965ed500540c189787b571c Mon Sep 17 00:00:00 2001 From: Mac Malone Date: Wed, 4 Dec 2024 17:40:14 +0000 Subject: [PATCH 510/829] chore(CI): rm duplicate ProofWidgets step (#19729) Removes the ProofWidgets fetch step from CI. This is unnecessary as this is now done as part of `lake exe cache get` (a detail I failed to realize in #19524). --- .github/build.in.yml | 10 ---------- .github/workflows/bors.yml | 10 ---------- .github/workflows/build.yml | 10 ---------- .github/workflows/build_fork.yml | 10 ---------- 4 files changed, 40 deletions(-) diff --git a/.github/build.in.yml b/.github/build.in.yml index 141b5316926ab..6a1d17a8cf495 100644 --- a/.github/build.in.yml +++ b/.github/build.in.yml @@ -127,16 +127,6 @@ jobs: run: | lake build cache - - name: prune ProofWidgets .lake - run: | - lake build proofwidgets:release - # The ProofWidgets release contains not just the `.js` (which we need in order to build) - # but also `.oleans`, which may have been built with the wrong toolchain. - # This removes them. - # See discussion at https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/nightly-testing/near/411225235 - rm -rf .lake/packages/proofwidgets/.lake/build/lib - rm -rf .lake/packages/proofwidgets/.lake/build/ir - - name: get cache id: get run: | diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index 3bd02efbb94dd..87e3093f9b131 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -137,16 +137,6 @@ jobs: run: | lake build cache - - name: prune ProofWidgets .lake - run: | - lake build proofwidgets:release - # The ProofWidgets release contains not just the `.js` (which we need in order to build) - # but also `.oleans`, which may have been built with the wrong toolchain. - # This removes them. - # See discussion at https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/nightly-testing/near/411225235 - rm -rf .lake/packages/proofwidgets/.lake/build/lib - rm -rf .lake/packages/proofwidgets/.lake/build/ir - - name: get cache id: get run: | diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index b5f11a1ea1f03..5ae1b33476e94 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -144,16 +144,6 @@ jobs: run: | lake build cache - - name: prune ProofWidgets .lake - run: | - lake build proofwidgets:release - # The ProofWidgets release contains not just the `.js` (which we need in order to build) - # but also `.oleans`, which may have been built with the wrong toolchain. - # This removes them. - # See discussion at https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/nightly-testing/near/411225235 - rm -rf .lake/packages/proofwidgets/.lake/build/lib - rm -rf .lake/packages/proofwidgets/.lake/build/ir - - name: get cache id: get run: | diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index 4da591010b8b7..79bfe1d0f3d0a 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -141,16 +141,6 @@ jobs: run: | lake build cache - - name: prune ProofWidgets .lake - run: | - lake build proofwidgets:release - # The ProofWidgets release contains not just the `.js` (which we need in order to build) - # but also `.oleans`, which may have been built with the wrong toolchain. - # This removes them. - # See discussion at https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/nightly-testing/near/411225235 - rm -rf .lake/packages/proofwidgets/.lake/build/lib - rm -rf .lake/packages/proofwidgets/.lake/build/ir - - name: get cache id: get run: | From 0257491fbdb03fe780243254ddb41ffd82f53385 Mon Sep 17 00:00:00 2001 From: Antoine Chambert-Loir Date: Wed, 4 Dec 2024 18:37:22 +0000 Subject: [PATCH 511/829] feat(Data/Finsupp/MonomialOrder/DegLex): homogeneous lexicographic order (#19455) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Definition of the homogeneous lexicographic order This is part of an ongoing work on basic of Gröbner theory. A subsequent PR will add the homogeneous reverse lexicographic order. --- Mathlib.lean | 1 + Mathlib/Data/Finsupp/Lex.lean | 20 ++ .../Data/Finsupp/MonomialOrder/DegLex.lean | 258 ++++++++++++++++++ Mathlib/Data/Finsupp/Weight.lean | 14 + 4 files changed, 293 insertions(+) create mode 100644 Mathlib/Data/Finsupp/MonomialOrder/DegLex.lean diff --git a/Mathlib.lean b/Mathlib.lean index 144457d89beae..55139cfd2afd8 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2483,6 +2483,7 @@ import Mathlib.Data.Finsupp.Indicator import Mathlib.Data.Finsupp.Interval import Mathlib.Data.Finsupp.Lex import Mathlib.Data.Finsupp.MonomialOrder +import Mathlib.Data.Finsupp.MonomialOrder.DegLex import Mathlib.Data.Finsupp.Multiset import Mathlib.Data.Finsupp.NeLocus import Mathlib.Data.Finsupp.Notation diff --git a/Mathlib/Data/Finsupp/Lex.lean b/Mathlib/Data/Finsupp/Lex.lean index 6b2a45dd8483b..87e567ec57773 100644 --- a/Mathlib/Data/Finsupp/Lex.lean +++ b/Mathlib/Data/Finsupp/Lex.lean @@ -75,6 +75,26 @@ instance Lex.linearOrder [LinearOrder N] : LinearOrder (Lex (α →₀ N)) where le := (· ≤ ·) __ := LinearOrder.lift' (toLex ∘ toDFinsupp ∘ ofLex) finsuppEquivDFinsupp.injective +theorem Lex.single_strictAnti : StrictAnti (fun (a : α) ↦ toLex (single a 1)) := by + intro a b h + simp only [LT.lt, Finsupp.lex_def] + simp only [ofLex_toLex, Nat.lt_eq] + use a + constructor + · intro d hd + simp only [Finsupp.single_eq_of_ne (ne_of_lt hd).symm, + Finsupp.single_eq_of_ne (ne_of_gt (lt_trans hd h))] + · simp only [single_eq_same, single_eq_of_ne (ne_of_lt h).symm, zero_lt_one] + +theorem Lex.single_lt_iff {a b : α} : toLex (single b 1) < toLex (single a 1) ↔ a < b := + Lex.single_strictAnti.lt_iff_lt + +theorem Lex.single_le_iff {a b : α} : toLex (single b 1) ≤ toLex (single a 1) ↔ a ≤ b := + Lex.single_strictAnti.le_iff_le + +theorem Lex.single_antitone : Antitone (fun (a : α) ↦ toLex (single a 1)) := + Lex.single_strictAnti.antitone + variable [PartialOrder N] theorem toLex_monotone : Monotone (@toLex (α →₀ N)) := diff --git a/Mathlib/Data/Finsupp/MonomialOrder/DegLex.lean b/Mathlib/Data/Finsupp/MonomialOrder/DegLex.lean new file mode 100644 index 0000000000000..0deee8bfefe7c --- /dev/null +++ b/Mathlib/Data/Finsupp/MonomialOrder/DegLex.lean @@ -0,0 +1,258 @@ +/- +Copyright (c) 2024 Antoine Chambert-Loir. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Antoine Chambert-Loir +-/ + +import Mathlib.Data.Finsupp.MonomialOrder +import Mathlib.Data.Finsupp.Weight +import Mathlib.Logic.Equiv.TransferInstance + +/-! Homogeneous lexicographic monomial ordering + +* `MonomialOrder.degLex`: a variant of the lexicographic ordering that first compares degrees. +For this, `σ` needs to be embedded with an ordering relation which satisfies `WellFoundedGT σ`. +(This last property is automatic when `σ` is finite). + +The type synonym is `DegLex (σ →₀ ℕ)` and the two lemmas `MonomialOrder.degLex_le_iff` +and `MonomialOrder.degLex_lt_iff` rewrite the ordering as comparisons in the type `Lex (σ →₀ ℕ)`. + +## References + +* [Cox, Little and O'Shea, *Ideals, varieties, and algorithms*][coxlittleoshea1997] +* [Becker and Weispfenning, *Gröbner bases*][Becker-Weispfenning1993] + +-/ + +section degLex + +/-- A type synonym to equip a type with its lexicographic order sorted by degrees. -/ +def DegLex (α : Type*) := α + +variable {α : Type*} + +/-- `toDegLex` is the identity function to the `DegLex` of a type. -/ +@[match_pattern] def toDegLex : α ≃ DegLex α := Equiv.refl _ + +theorem toDegLex_injective : Function.Injective (toDegLex (α := α)) := fun _ _ ↦ _root_.id + +theorem toDegLex_inj {a b : α} : toDegLex a = toDegLex b ↔ a = b := Iff.rfl + +/-- `ofDegLex` is the identity function from the `DegLex` of a type. -/ +@[match_pattern] def ofDegLex : DegLex α ≃ α := Equiv.refl _ + +theorem ofDegLex_injective : Function.Injective (ofDegLex (α := α)) := fun _ _ ↦ _root_.id + +theorem ofDegLex_inj {a b : DegLex α} : ofDegLex a = ofDegLex b ↔ a = b := Iff.rfl + +@[simp] theorem ofDegLex_symm_eq : (@ofDegLex α).symm = toDegLex := rfl + +@[simp] theorem toDegLex_symm_eq : (@toDegLex α).symm = ofDegLex := rfl + +@[simp] theorem ofDegLex_toDegLex (a : α) : ofDegLex (toDegLex a) = a := rfl + +@[simp] theorem toDegLex_ofDegLex (a : DegLex α) : toDegLex (ofDegLex a) = a := rfl + +/-- A recursor for `DegLex`. Use as `induction x`. -/ +@[elab_as_elim, induction_eliminator, cases_eliminator] +protected def DegLex.rec {β : DegLex α → Sort*} (h : ∀ a, β (toDegLex a)) : + ∀ a, β a := fun a => h (ofDegLex a) + +@[simp] lemma DegLex.forall_iff {p : DegLex α → Prop} : (∀ a, p a) ↔ ∀ a, p (toDegLex a) := Iff.rfl +@[simp] lemma DegLex.exists_iff {p : DegLex α → Prop} : (∃ a, p a) ↔ ∃ a, p (toDegLex a) := Iff.rfl + +noncomputable instance [AddCommMonoid α] : + AddCommMonoid (DegLex α) := ofDegLex.addCommMonoid + +theorem toDegLex_add [AddCommMonoid α] (a b : α) : + toDegLex (a + b) = toDegLex a + toDegLex b := rfl + +theorem ofDegLex_add [AddCommMonoid α] (a b : DegLex α) : + ofDegLex (a + b) = ofDegLex a + ofDegLex b := rfl + +namespace Finsupp + +/-- `Finsupp.DegLex r s` is the homogeneous lexicographic order on `α →₀ M`, +where `α` is ordered by `r` and `M` is ordered by `s`. +The type synonym `DegLex (α →₀ M)` has an order given by `Finsupp.DegLex (· < ·) (· < ·)`. -/ +protected def DegLex (r : α → α → Prop) (s : ℕ → ℕ → Prop) : + (α →₀ ℕ) → (α →₀ ℕ) → Prop := + (Prod.Lex s (Finsupp.Lex r s)) on (fun x ↦ (x.degree, x)) + +theorem degLex_def {r : α → α → Prop} {s : ℕ → ℕ → Prop} {a b : α →₀ ℕ} : + Finsupp.DegLex r s a b ↔ Prod.Lex s (Finsupp.Lex r s) (a.degree, a) (b.degree, b) := + Iff.rfl + +namespace DegLex + +theorem wellFounded + {r : α → α → Prop} [IsTrichotomous α r] (hr : WellFounded (Function.swap r)) + {s : ℕ → ℕ → Prop} (hs : WellFounded s) (hs0 : ∀ ⦃n⦄, ¬ s n 0) : + WellFounded (Finsupp.DegLex r s) := by + have wft := WellFounded.prod_lex hs (Finsupp.Lex.wellFounded' hs0 hs hr) + rw [← Set.wellFoundedOn_univ] at wft + unfold Finsupp.DegLex + rw [← Set.wellFoundedOn_range] + exact Set.WellFoundedOn.mono wft (le_refl _) (fun _ _ ↦ trivial) + +instance [LT α] : LT (DegLex (α →₀ ℕ)) := + ⟨fun f g ↦ Finsupp.DegLex (· < ·) (· < ·) (ofDegLex f) (ofDegLex g)⟩ + +theorem lt_def [LT α] {a b : DegLex (α →₀ ℕ)} : + a < b ↔ (toLex ((ofDegLex a).degree, toLex (ofDegLex a))) < + (toLex ((ofDegLex b).degree, toLex (ofDegLex b))) := + Iff.rfl + +theorem lt_iff [LT α] {a b : DegLex (α →₀ ℕ)} : + a < b ↔ (ofDegLex a).degree < (ofDegLex b).degree ∨ + (((ofDegLex a).degree = (ofDegLex b).degree) ∧ toLex (ofDegLex a) < toLex (ofDegLex b)) := by + simp only [lt_def, Prod.Lex.lt_iff] + +variable [LinearOrder α] + +instance isStrictOrder : IsStrictOrder (DegLex (α →₀ ℕ)) (· < ·) where + irrefl := fun a ↦ by simp [lt_def] + trans := by + intro a b c hab hbc + simp only [lt_iff] at hab hbc ⊢ + rcases hab with (hab | hab) + · rcases hbc with (hbc | hbc) + · left; exact lt_trans hab hbc + · left; exact lt_of_lt_of_eq hab hbc.1 + · rcases hbc with (hbc | hbc) + · left; exact lt_of_eq_of_lt hab.1 hbc + · right; exact ⟨Eq.trans hab.1 hbc.1, lt_trans hab.2 hbc.2⟩ + +/-- The linear order on `Finsupp`s obtained by the homogeneous lexicographic ordering. -/ +instance : LinearOrder (DegLex (α →₀ ℕ)) := + LinearOrder.lift' + (fun (f : DegLex (α →₀ ℕ)) ↦ toLex ((ofDegLex f).degree, toLex (ofDegLex f))) + (fun f g ↦ by simp) + +theorem le_iff {x y : DegLex (α →₀ ℕ)} : + x ≤ y ↔ (ofDegLex x).degree < (ofDegLex y).degree ∨ + (ofDegLex x).degree = (ofDegLex y).degree ∧ toLex (ofDegLex x) ≤ toLex (ofDegLex y) := by + simp only [le_iff_eq_or_lt, lt_iff, EmbeddingLike.apply_eq_iff_eq] + by_cases h : x = y + · simp [h] + · by_cases k : (ofDegLex x).degree < (ofDegLex y).degree + · simp [k] + · simp only [h, k, false_or] + +noncomputable instance : OrderedCancelAddCommMonoid (DegLex (α →₀ ℕ)) where + le_of_add_le_add_left a b c h := by + rw [le_iff] at h ⊢ + simpa only [ofDegLex_add, degree_add, add_lt_add_iff_left, add_right_inj, toLex_add, + add_le_add_iff_left] using h + add_le_add_left a b h c := by + rw [le_iff] at h ⊢ + simpa [ofDegLex_add, degree_add] using h + +/-- The linear order on `Finsupp`s obtained by the homogeneous lexicographic ordering. -/ +noncomputable instance : + LinearOrderedCancelAddCommMonoid (DegLex (α →₀ ℕ)) where + le_total := instLinearOrderDegLexNat.le_total + decidableLE := instLinearOrderDegLexNat.decidableLE + compare_eq_compareOfLessAndEq := instLinearOrderDegLexNat.compare_eq_compareOfLessAndEq + +theorem single_strictAnti : StrictAnti (fun (a : α) ↦ toDegLex (single a 1)) := by + intro _ _ h + simp only [lt_iff, ofDegLex_toDegLex, degree_single, lt_self_iff_false, Lex.single_lt_iff, h, + and_self, or_true] + +theorem single_antitone : Antitone (fun (a : α) ↦ toDegLex (single a 1)) := + single_strictAnti.antitone + +theorem single_lt_iff {a b : α} : + toDegLex (Finsupp.single b 1) < toDegLex (Finsupp.single a 1) ↔ a < b := + single_strictAnti.lt_iff_lt + +theorem single_le_iff {a b : α} : + toDegLex (Finsupp.single b 1) ≤ toDegLex (Finsupp.single a 1) ↔ a ≤ b := + single_strictAnti.le_iff_le + +theorem monotone_degree : + Monotone (fun (x : DegLex (α →₀ ℕ)) ↦ (ofDegLex x).degree) := by + intro x y + rw [le_iff] + rintro (h | h) + · apply le_of_lt h + · apply le_of_eq h.1 + +instance orderBot : OrderBot (DegLex (α →₀ ℕ)) where + bot := toDegLex (0 : α →₀ ℕ) + bot_le x := by + simp only [le_iff, ofDegLex_toDegLex, toLex_zero, degree_zero] + rcases eq_zero_or_pos (ofDegLex x).degree with (h | h) + · simp only [h, lt_self_iff_false, true_and, false_or, ge_iff_le] + exact bot_le + · simp [h] + +instance wellFoundedLT [WellFoundedGT α] : + WellFoundedLT (DegLex (α →₀ ℕ)) := + ⟨wellFounded wellFounded_gt wellFounded_lt fun n ↦ (zero_le n).not_lt⟩ + +end DegLex + +end Finsupp + +namespace MonomialOrder + +open Finsupp + +variable {σ : Type*} [LinearOrder σ] [WellFoundedGT σ] + +/-- The deg-lexicographic order on `σ →₀ ℕ`, as a `MonomialOrder` -/ +noncomputable def degLex : + MonomialOrder σ where + syn := DegLex (σ →₀ ℕ) + toSyn := { toEquiv := toDegLex, map_add' := toDegLex_add } + toSyn_monotone a b h := by + simp only [AddEquiv.coe_mk, DegLex.le_iff, ofDegLex_toDegLex] + by_cases ha : a.degree < b.degree + · exact Or.inl ha + · refine Or.inr ⟨le_antisymm ?_ (not_lt.mp ha), toLex_monotone h⟩ + rw [← add_tsub_cancel_of_le h, degree_add] + exact Nat.le_add_right a.degree (b - a).degree + +theorem degLex_le_iff {a b : σ →₀ ℕ} : + a ≼[degLex] b ↔ toDegLex a ≤ toDegLex b := + Iff.rfl + +theorem degLex_lt_iff {a b : σ →₀ ℕ} : + a ≺[degLex] b ↔ toDegLex a < toDegLex b := + Iff.rfl + +theorem degLex_single_le_iff {a b : σ} : + single a 1 ≼[degLex] single b 1 ↔ b ≤ a := by + rw [MonomialOrder.degLex_le_iff, DegLex.single_le_iff] + +theorem degLex_single_lt_iff {a b : σ} : + single a 1 ≺[degLex] single b 1 ↔ b < a := by + rw [MonomialOrder.degLex_lt_iff, DegLex.single_lt_iff] + +end MonomialOrder + +section Examples + +open Finsupp MonomialOrder DegLex + +/-- for the deg-lexicographic ordering, X 1 < X 0 -/ +example : single (1 : Fin 2) 1 ≺[degLex] single 0 1 := by + rw [degLex_lt_iff, single_lt_iff] + exact Nat.one_pos + +/-- for the deg-lexicographic ordering, X 0 * X 1 < X 0 ^ 2 -/ +example : (single 0 1 + single 1 1) ≺[degLex] single (0 : Fin 2) 2 := by + simp only [degLex_lt_iff, lt_iff, ofDegLex_toDegLex, degree_add, + degree_single, Nat.reduceAdd, lt_self_iff_false, true_and, false_or] + use 0 + simp + +/-- for the deg-lexicographic ordering, X 0 < X 1 ^ 2 -/ +example : single (0 : Fin 2) 1 ≺[degLex] single 1 2 := by + simp [degLex_lt_iff, lt_iff] + +end Examples + +end degLex diff --git a/Mathlib/Data/Finsupp/Weight.lean b/Mathlib/Data/Finsupp/Weight.lean index 7b07c03ebf167..36ec2f64e9d94 100644 --- a/Mathlib/Data/Finsupp/Weight.lean +++ b/Mathlib/Data/Finsupp/Weight.lean @@ -202,6 +202,20 @@ def degree (d : σ →₀ ℕ) := ∑ i ∈ d.support, d i @[deprecated degree (since := "2024-07-20")] alias _root_.MvPolynomial.degree := degree +@[simp] +theorem degree_add (a b : σ →₀ ℕ) : (a + b).degree = a.degree + b.degree := + sum_add_index' (h := fun _ ↦ id) (congrFun rfl) fun _ _ ↦ congrFun rfl + +@[simp] +theorem degree_single (a : σ) (m : ℕ) : (Finsupp.single a m).degree = m := by + rw [degree, Finset.sum_eq_single a] + · simp only [single_eq_same] + · intro b _ hba + exact single_eq_of_ne hba.symm + · intro ha + simp only [mem_support_iff, single_eq_same, ne_eq, Decidable.not_not] at ha + rw [single_eq_same, ha] + lemma degree_eq_zero_iff (d : σ →₀ ℕ) : degree d = 0 ↔ d = 0 := by simp only [degree, Finset.sum_eq_zero_iff, Finsupp.mem_support_iff, ne_eq, Decidable.not_imp_self, DFunLike.ext_iff, Finsupp.coe_zero, Pi.zero_apply] From 58c96928b554e7a21f27914368edae0a2167a608 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Wed, 4 Dec 2024 19:43:06 +0000 Subject: [PATCH 512/829] feat(LinearAlgebra/ExteriorPower): the universal property of the exterior power (#18590) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We obtain the universal property of the `n`th exterior power of a module `M`: linear maps from this exterior power identify to `n`-alternating maps from `M`. This is also phrased as the data of a presentation of the exterior power by generators and relations. Co-authored-by: sophie.morel@ens-lyon.fr Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- Mathlib.lean | 1 + .../LinearAlgebra/ExteriorPower/Basic.lean | 201 ++++++++++++++++++ Mathlib/LinearAlgebra/Span/Basic.lean | 51 +++++ 3 files changed, 253 insertions(+) create mode 100644 Mathlib/LinearAlgebra/ExteriorPower/Basic.lean diff --git a/Mathlib.lean b/Mathlib.lean index 55139cfd2afd8..3a5b5b5194894 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3334,6 +3334,7 @@ import Mathlib.LinearAlgebra.Eigenspace.Zero import Mathlib.LinearAlgebra.ExteriorAlgebra.Basic import Mathlib.LinearAlgebra.ExteriorAlgebra.Grading import Mathlib.LinearAlgebra.ExteriorAlgebra.OfAlternating +import Mathlib.LinearAlgebra.ExteriorPower.Basic import Mathlib.LinearAlgebra.FiniteDimensional import Mathlib.LinearAlgebra.FiniteDimensional.Defs import Mathlib.LinearAlgebra.FiniteSpan diff --git a/Mathlib/LinearAlgebra/ExteriorPower/Basic.lean b/Mathlib/LinearAlgebra/ExteriorPower/Basic.lean new file mode 100644 index 0000000000000..e02ccb29fa8ea --- /dev/null +++ b/Mathlib/LinearAlgebra/ExteriorPower/Basic.lean @@ -0,0 +1,201 @@ +/- +Copyright (c) 2024 Sophie Morel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sophie Morel, Joël Riou +-/ +import Mathlib.Algebra.Module.Presentation.Basic +import Mathlib.LinearAlgebra.ExteriorAlgebra.OfAlternating + +/-! +# Exterior powers + +We study the exterior powers of a module `M` over a commutative ring `R`. + +## Definitions + +* `exteriorPower.ιMulti` is the canonical alternating map on `M` with values in `⋀[R]^n M`. + +* `exteriorPower.presentation R n M` is the standard presentation of the `R`-module `⋀[R]^n M`. + +## Theorems +* `ιMulti_span`: The image of `exteriorPower.ιMulti` spans `⋀[R]^n M`. + +* We construct `exteriorPower.alternatingMapLinearEquiv` which +expresses the universal property of the exterior power as a +linear equivalence `(M [⋀^Fin n]→ₗ[R] N) ≃ₗ[R] ⋀[R]^n M →ₗ[R] N` between +alternating maps and linear maps from the exterior power. + +-/ + +open scoped TensorProduct + +universe u v uM uN uN' uN'' uE uF + +variable (R : Type u) [CommRing R] (n : ℕ) {M N N' : Type*} + [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] + [AddCommGroup N'] [Module R N'] + +namespace exteriorPower + +open Function + +/-! The canonical alternating map from `Fin n → M` to `⋀[R]^n M`. -/ + +/-- `exteriorAlgebra.ιMulti` is the alternating map from `Fin n → M` to `⋀[r]^n M` +induced by `exteriorAlgebra.ιMulti`, i.e. sending a family of vectors `m : Fin n → M` to the +product of its entries. -/ +def ιMulti : M [⋀^Fin n]→ₗ[R] (⋀[R]^n M) := + (ExteriorAlgebra.ιMulti R n).codRestrict (⋀[R]^n M) fun _ => + ExteriorAlgebra.ιMulti_range R n <| Set.mem_range_self _ + +@[simp] lemma ιMulti_apply_coe (a : Fin n → M) : ιMulti R n a = ExteriorAlgebra.ιMulti R n a := rfl + +variable (M) +/-- The image of `ExteriorAlgebra.ιMulti R n` spans the `n`th exterior power. Variant of +`ExteriorAlgebra.ιMulti_span_fixedDegree`, useful in rewrites. -/ +lemma ιMulti_span_fixedDegree : + Submodule.span R (Set.range (ExteriorAlgebra.ιMulti R n)) = ⋀[R]^n M := + ExteriorAlgebra.ιMulti_span_fixedDegree R n + +/-- The image of `exteriorPower.ιMulti` spans `⋀[R]^n M`. -/ +lemma ιMulti_span : + Submodule.span R (Set.range (ιMulti R n)) = (⊤ : Submodule R (⋀[R]^n M)) := by + apply LinearMap.map_injective (Submodule.ker_subtype (⋀[R]^n M)) + rw [LinearMap.map_span, ← Set.image_univ, Set.image_image] + simp only [Submodule.coe_subtype, ιMulti_apply_coe, Set.image_univ, Submodule.map_top, + Submodule.range_subtype] + exact ExteriorAlgebra.ιMulti_span_fixedDegree R n + +namespace presentation + +/-- The index type for the relations in the standard presentation of `⋀[R]^n M`, +in the particular case `ι` is `Fin n`. -/ +inductive Rels (ι : Type*) (M : Type*) + | add (m : ι → M) (i : ι) (x y : M) + | smul (m : ι → M) (i : ι) (r : R) (x : M) + | alt (m : ι → M) (i j : ι) (hm : m i = m j) (hij : i ≠ j) + +/-- The relations in the standard presentation of `⋀[R]^n M` with generators and relations. -/ +@[simps] +noncomputable def relations (ι : Type*) [DecidableEq ι] (M : Type*) + [AddCommGroup M] [Module R M] : + Module.Relations R where + G := ι → M + R := Rels R ι M + relation r := match r with + | .add m i x y => Finsupp.single (update m i x) 1 + + Finsupp.single (update m i y) 1 - + Finsupp.single (update m i (x + y)) 1 + | .smul m i r x => Finsupp.single (update m i (r • x)) 1 - + r • Finsupp.single (update m i x) 1 + | .alt m _ _ _ _ => Finsupp.single m 1 + +variable {R} in +/-- The solutions in a module `N` to the linear equations +given by `exteriorPower.relations R ι M` identify to alternating maps to `N`. -/ +@[simps!] +def relationsSolutionEquiv {ι : Type*} [DecidableEq ι] {M : Type*} + [AddCommGroup M] [Module R M] : + (relations R ι M).Solution N ≃ AlternatingMap R M N ι where + toFun s := + { toFun := fun m ↦ s.var m + map_update_add' := fun m i x y ↦ by + have := s.linearCombination_var_relation (.add m i x y) + dsimp at this ⊢ + rw [map_sub, map_add, Finsupp.linearCombination_single, one_smul, + Finsupp.linearCombination_single, one_smul, + Finsupp.linearCombination_single, one_smul, sub_eq_zero] at this + convert this.symm -- `convert` is necessary due to the implementation of `MultilinearMap` + map_update_smul' := fun m i r x ↦ by + have := s.linearCombination_var_relation (.smul m i r x) + dsimp at this ⊢ + rw [Finsupp.smul_single, smul_eq_mul, mul_one, map_sub, + Finsupp.linearCombination_single, one_smul, + Finsupp.linearCombination_single, sub_eq_zero] at this + convert this + map_eq_zero_of_eq' := fun v i j hm hij ↦ + by simpa using s.linearCombination_var_relation (.alt v i j hm hij) } + invFun f := + { var := fun m ↦ f m + linearCombination_var_relation := by + rintro (⟨m, i, x, y⟩ | ⟨m, i, r, x⟩ | ⟨v, i, j, hm, hij⟩) + · simp + · simp + · simpa using f.map_eq_zero_of_eq v hm hij } + left_inv _ := rfl + right_inv _ := rfl + +/-- The universal property of the exterior power. -/ +def isPresentationCore : + (relationsSolutionEquiv.symm (ιMulti R n (M := M))).IsPresentationCore where + desc s := LinearMap.comp (ExteriorAlgebra.liftAlternating + (Function.update 0 n (relationsSolutionEquiv s))) (Submodule.subtype _) + postcomp_desc s := by aesop + postcomp_injective {N _ _ f f' h} := by + rw [Submodule.linearMap_eq_iff_of_span_eq_top _ _ (ιMulti_span R n M)] + rintro ⟨_, ⟨f, rfl⟩⟩ + exact Module.Relations.Solution.congr_var h f + +end presentation + +/-- The standard presentation of the `R`-module `⋀[R]^n M`. -/ +@[simps! G R relation var] +noncomputable def presentation : Module.Presentation R (⋀[R]^n M) := + .ofIsPresentation (presentation.isPresentationCore R n M).isPresentation + +variable {R M n} + +/-- Two linear maps on `⋀[R]^n M` that agree on the image of `exteriorPower.ιMulti` +are equal. -/ +@[ext] +lemma linearMap_ext {f : ⋀[R]^n M →ₗ[R] N} {g : ⋀[R]^n M →ₗ[R] N} + (heq : f.compAlternatingMap (ιMulti R n) = g.compAlternatingMap (ιMulti R n)) : f = g := + (presentation R n M).postcomp_injective (by ext f; apply DFunLike.congr_fun heq ) + +/-- The linear equivalence between `n`-fold alternating maps from `M` to `N` and linear maps from +`⋀[R]^n M` to `N`: this is the universal property of the `n`th exterior power of `M`. -/ +noncomputable def alternatingMapLinearEquiv : (M [⋀^Fin n]→ₗ[R] N) ≃ₗ[R] ⋀[R]^n M →ₗ[R] N := + LinearEquiv.symm + (Equiv.toLinearEquiv + ((presentation R n M).linearMapEquiv.trans presentation.relationsSolutionEquiv) + { map_add := fun _ _ => rfl + map_smul := fun _ _ => rfl }) + +@[simp] +lemma alternatingMapLinearEquiv_comp_ιMulti (f : M [⋀^Fin n]→ₗ[R] N) : + (alternatingMapLinearEquiv f).compAlternatingMap (ιMulti R n) = f := by + obtain ⟨φ, rfl⟩ := alternatingMapLinearEquiv.symm.surjective f + dsimp [alternatingMapLinearEquiv] + simp only [LinearEquiv.symm_apply_apply] + rfl + +@[simp] +lemma alternatingMapLinearEquiv_apply_ιMulti (f : M [⋀^Fin n]→ₗ[R] N) (a : Fin n → M) : + alternatingMapLinearEquiv f (ιMulti R n a) = f a := + DFunLike.congr_fun (alternatingMapLinearEquiv_comp_ιMulti f) a + +@[simp] +lemma alternatingMapLinearEquiv_symm_apply (F : ⋀[R]^n M →ₗ[R] N) (m : Fin n → M) : + alternatingMapLinearEquiv.symm F m = F.compAlternatingMap (ιMulti R n) m := by + obtain ⟨f, rfl⟩ := alternatingMapLinearEquiv.surjective F + simp only [LinearEquiv.symm_apply_apply, alternatingMapLinearEquiv_comp_ιMulti] + +@[simp] +lemma alternatingMapLinearEquiv_ιMulti : + alternatingMapLinearEquiv (ιMulti R n (M := M)) = LinearMap.id := by + ext + simp only [alternatingMapLinearEquiv_comp_ιMulti, ιMulti_apply_coe, + LinearMap.compAlternatingMap_apply, LinearMap.id_coe, id_eq] + +/-- If `f` is an alternating map from `M` to `N`, +`alternatingMapLinearEquiv f` is the corresponding linear map from `⋀[R]^n M` to `N`, +and if `g` is a linear map from `N` to `N'`, then +the alternating map `g.compAlternatingMap f` from `M` to `N'` corresponds to the linear +map `g.comp (alternatingMapLinearEquiv f)` on `⋀[R]^n M`. -/ +lemma alternatingMapLinearEquiv_comp (g : N →ₗ[R] N') (f : M [⋀^Fin n]→ₗ[R] N) : + alternatingMapLinearEquiv (g.compAlternatingMap f) = g.comp (alternatingMapLinearEquiv f) := by + ext + simp only [alternatingMapLinearEquiv_comp_ιMulti, LinearMap.compAlternatingMap_apply, + LinearMap.coe_comp, comp_apply, alternatingMapLinearEquiv_apply_ιMulti] + +end exteriorPower diff --git a/Mathlib/LinearAlgebra/Span/Basic.lean b/Mathlib/LinearAlgebra/Span/Basic.lean index 0a1c4349e44cc..997c8ed3189d4 100644 --- a/Mathlib/LinearAlgebra/Span/Basic.lean +++ b/Mathlib/LinearAlgebra/Span/Basic.lean @@ -80,6 +80,57 @@ theorem span_preimage_le (f : F) (s : Set M₂) : alias _root_.LinearMap.span_preimage_le := Submodule.span_preimage_le +section + +variable {N : Type*} [AddCommMonoid N] [Module R N] + +lemma linearMap_eq_iff_of_eq_span {V : Submodule R M} (f g : V →ₗ[R] N) + {S : Set M} (hV : V = span R S) : + f = g ↔ ∀ (s : S), f ⟨s, by simpa only [hV] using subset_span (by simp)⟩ = + g ⟨s, by simpa only [hV] using subset_span (by simp)⟩ := by + constructor + · rintro rfl _ + rfl + · intro h + subst hV + suffices ∀ (x : M) (hx : x ∈ span R S), f ⟨x, hx⟩ = g ⟨x, hx⟩ by + ext ⟨x, hx⟩ + exact this x hx + intro x hx + induction hx using span_induction with + | mem x hx => exact h ⟨x, hx⟩ + | zero => erw [map_zero, map_zero] + | add x y hx hy hx' hy' => + erw [f.map_add ⟨x, hx⟩ ⟨y, hy⟩, g.map_add ⟨x, hx⟩ ⟨y, hy⟩] + rw [hx', hy'] + | smul a x hx hx' => + erw [f.map_smul a ⟨x, hx⟩, g.map_smul a ⟨x, hx⟩] + rw [hx'] + +lemma linearMap_eq_iff_of_span_eq_top (f g : M →ₗ[R] N) + {S : Set M} (hM : span R S = ⊤) : + f = g ↔ ∀ (s : S), f s = g s := by + convert linearMap_eq_iff_of_eq_span (f.comp (Submodule.subtype _)) + (g.comp (Submodule.subtype _)) hM.symm + constructor + · rintro rfl + rfl + · intro h + ext x + exact DFunLike.congr_fun h ⟨x, by simp⟩ + +lemma linearMap_eq_zero_iff_of_span_eq_top (f : M →ₗ[R] N) + {S : Set M} (hM : span R S = ⊤) : + f = 0 ↔ ∀ (s : S), f s = 0 := + linearMap_eq_iff_of_span_eq_top f 0 hM + +lemma linearMap_eq_zero_iff_of_eq_span {V : Submodule R M} (f : V →ₗ[R] N) + {S : Set M} (hV : V = span R S) : + f = 0 ↔ ∀ (s : S), f ⟨s, by simpa only [hV] using subset_span (by simp)⟩ = 0 := + linearMap_eq_iff_of_eq_span f 0 hV + +end + /-- See `Submodule.span_smul_eq` (in `RingTheory.Ideal.Operations`) for `span R (r • s) = r • span R s` that holds for arbitrary `r` in a `CommSemiring`. -/ theorem span_smul_eq_of_isUnit (s : Set M) (r : R) (hr : IsUnit r) : span R (r • s) = span R s := by From fa8f17575769f525aafdcddbec606c939e7c0363 Mon Sep 17 00:00:00 2001 From: Emily Riehl Date: Wed, 4 Dec 2024 20:12:49 +0000 Subject: [PATCH 513/829] feat(AlgebraicTopology/SimplicialSet): StrictSegal simplicial sets are 2-coskeletal (#19057) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We prove that if `X` is a simplicial set that is `StrictSegal` then `X` is 2-coskeletal, meaning `(coskAdj 2).unit.app X` defines an isomorphism `StrictSegal.cosk2Iso : X ≅ (Truncated.cosk 2).obj ((truncation 2).obj X)`. This is proven by demonstrating that `(rightExtensionInclusion X n).IsPointwiseRightKanExtension` holds. Co-Authored-By: [Mario Carneiro](https://github.com/digama0) and [Joël Riou](https://github.com/joelriou) Co-authored-by: Joël Riou --- Mathlib.lean | 1 + .../AlgebraicTopology/SimplexCategory.lean | 4 + .../SimplicialSet/Basic.lean | 4 +- .../SimplicialSet/Coskeletal.lean | 256 ++++++++++++++++++ .../SimplicialSet/Nerve.lean | 2 +- .../AlgebraicTopology/SimplicialSet/Path.lean | 7 + .../SimplicialSet/StrictSegal.lean | 4 + 7 files changed, 275 insertions(+), 3 deletions(-) create mode 100644 Mathlib/AlgebraicTopology/SimplicialSet/Coskeletal.lean diff --git a/Mathlib.lean b/Mathlib.lean index 3a5b5b5194894..a11769a77bada 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1055,6 +1055,7 @@ import Mathlib.AlgebraicTopology.SimplicialCategory.SimplicialObject import Mathlib.AlgebraicTopology.SimplicialObject.Basic import Mathlib.AlgebraicTopology.SimplicialObject.Coskeletal import Mathlib.AlgebraicTopology.SimplicialSet.Basic +import Mathlib.AlgebraicTopology.SimplicialSet.Coskeletal import Mathlib.AlgebraicTopology.SimplicialSet.KanComplex import Mathlib.AlgebraicTopology.SimplicialSet.Monoidal import Mathlib.AlgebraicTopology.SimplicialSet.Nerve diff --git a/Mathlib/AlgebraicTopology/SimplexCategory.lean b/Mathlib/AlgebraicTopology/SimplexCategory.lean index 115b2bc0a1851..156dc3423a3ca 100644 --- a/Mathlib/AlgebraicTopology/SimplexCategory.lean +++ b/Mathlib/AlgebraicTopology/SimplexCategory.lean @@ -231,6 +231,10 @@ def mkOfLe {n} (i j : Fin (n+1)) (h : i ≤ j) : ([1] : SimplexCategory) ⟶ [n] | 0, 1, _ => h } +@[simp] +lemma mkOfLe_refl {n} (j : Fin (n + 1)) : + mkOfLe j j (by omega) = [1].const [n] j := Hom.ext_one_left _ _ + /-- The morphism `[1] ⟶ [n]` that picks out the "diagonal composite" edge-/ def diag (n : ℕ) : ([1] : SimplexCategory) ⟶ [n] := mkOfLe 0 n (Fin.zero_le _) diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean index 6e545ab1fa48e..be2e7617b9773 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean @@ -384,10 +384,10 @@ protected abbrev Truncated.cosk (n : ℕ) : SSet.Truncated n ⥤ SSet.{u} := SimplicialObject.Truncated.cosk n /-- The n-skeleton as an endofunctor on `SSet`. -/ -abbrev sk (n : ℕ) : SSet ⥤ SSet := SimplicialObject.sk n +abbrev sk (n : ℕ) : SSet.{u} ⥤ SSet.{u} := SimplicialObject.sk n /-- The n-coskeleton as an endofunctor on `SSet`. -/ -abbrev cosk (n : ℕ) : SSet ⥤ SSet := SimplicialObject.cosk n +abbrev cosk (n : ℕ) : SSet.{u} ⥤ SSet.{u} := SimplicialObject.cosk n end diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Coskeletal.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Coskeletal.lean new file mode 100644 index 0000000000000..93eed2ed3a913 --- /dev/null +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Coskeletal.lean @@ -0,0 +1,256 @@ +/- +Copyright (c) 2024 Emily Riehl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro, Emily Riehl, Joël Riou +-/ +import Mathlib.AlgebraicTopology.SimplicialObject.Coskeletal +import Mathlib.AlgebraicTopology.SimplicialSet.StrictSegal +import Mathlib.CategoryTheory.Functor.KanExtension.Adjunction +import Mathlib.CategoryTheory.Functor.KanExtension.Basic + +/-! +# Coskeletal simplicial sets + +In this file, we prove that if `X` is `StrictSegal` then `X` is 2-coskeletal, +i.e. `X ≅ (cosk 2).obj X`. In particular, nerves of categories are 2-coskeletal. + +This isomorphism follows from the fact that `rightExtensionInclusion X 2` is a right Kan +extension. In fact, we show that when `X` is `StrictSegal` then +`(rightExtensionInclusion X n).IsPointwiseRightKanExtension` holds. + +As an example, `SimplicialObject.IsCoskeletal (nerve C) 2` shows that that nerves of categories +are 2-coskeletal. +-/ + + +universe v u + +open CategoryTheory Simplicial SimplexCategory Opposite Category Functor Limits + +namespace SSet + +namespace Truncated + +/-- The identity natural transformation exhibits a simplicial set as a right extension of its +restriction along `(Truncated.inclusion (n := n)).op`.-/ +@[simps!] +def rightExtensionInclusion (X : SSet.{u}) (n : ℕ) : + RightExtension (Truncated.inclusion (n := n)).op + ((Truncated.inclusion n).op ⋙ X) := RightExtension.mk _ (𝟙 _) + +end Truncated + +section + +local notation (priority := high) "[" n "]" => SimplexCategory.mk n + +local macro:max (priority := high) "[" n:term "]₂" : term => + `((⟨SimplexCategory.mk $n, by dsimp; omega⟩ : SimplexCategory.Truncated 2)) + +open StructuredArrow + +namespace StrictSegal +variable (X : SSet.{u}) [StrictSegal X] + +namespace isPointwiseRightKanExtensionAt + +/-- A morphism in `SimplexCategory` with domain `[0]`, `[1]`, or `[2]` defines an object in the +comma category `StructuredArrow (op [n]) (Truncated.inclusion (n := 2)).op`.-/ +abbrev strArrowMk₂ {i : ℕ} {n : ℕ} (φ : [i] ⟶ [n]) (hi : i ≤ 2) : + StructuredArrow (op [n]) (Truncated.inclusion (n := 2)).op := + StructuredArrow.mk (Y := op ⟨[i], hi⟩) (by exact φ.op) + +/-- Given a term in the cone over the diagram +`(proj (op [n]) ((Truncated.inclusion 2).op ⋙ (Truncated.inclusion 2).op ⋙ X)` where `X` is +Strict Segal, one can produce an `n`-simplex in `X`. -/ +@[simp] +noncomputable def lift {X : SSet.{u}} [StrictSegal X] {n} + (s : Cone (proj (op [n]) (Truncated.inclusion 2).op ⋙ + (Truncated.inclusion 2).op ⋙ X)) (x : s.pt) : X _[n] := + StrictSegal.spineToSimplex { + vertex := fun i ↦ s.π.app (.mk (Y := op [0]₂) (.op (SimplexCategory.const _ _ i))) x + arrow := fun i ↦ s.π.app (.mk (Y := op [1]₂) (.op (mkOfLe _ _ (Fin.castSucc_le_succ i)))) x + arrow_src := fun i ↦ by + let φ : strArrowMk₂ (mkOfLe _ _ (Fin.castSucc_le_succ i)) (by simp) ⟶ + strArrowMk₂ ([0].const _ i.castSucc) (by simp) := + StructuredArrow.homMk (δ 1).op + (Quiver.Hom.unop_inj (by ext x; fin_cases x; rfl)) + exact congr_fun (s.w φ) x + arrow_tgt := fun i ↦ by + dsimp + let φ : strArrowMk₂ (mkOfLe _ _ (Fin.castSucc_le_succ i)) (by simp) ⟶ + strArrowMk₂ ([0].const _ i.succ) (by simp) := + StructuredArrow.homMk (δ 0).op + (Quiver.Hom.unop_inj (by ext x; fin_cases x; rfl)) + exact congr_fun (s.w φ) x } + +lemma fac_aux₁ {n : ℕ} + (s : Cone (proj (op [n]) (Truncated.inclusion 2).op ⋙ (Truncated.inclusion 2).op ⋙ X)) + (x : s.pt) (i : ℕ) (hi : i < n) : + X.map (mkOfSucc ⟨i, hi⟩).op (lift s x) = + s.π.app (strArrowMk₂ (mkOfSucc ⟨i, hi⟩) (by omega)) x := by + dsimp [lift] + rw [spineToSimplex_arrow] + rfl + +lemma fac_aux₂ {n : ℕ} + (s : Cone (proj (op [n]) (Truncated.inclusion 2).op ⋙ (Truncated.inclusion 2).op ⋙ X)) + (x : s.pt) (i j : ℕ) (hij : i ≤ j) (hj : j ≤ n) : + X.map (mkOfLe ⟨i, by omega⟩ ⟨j, by omega⟩ hij).op (lift s x) = + s.π.app (strArrowMk₂ (mkOfLe ⟨i, by omega⟩ ⟨j, by omega⟩ hij) (by omega)) x := by + obtain ⟨k, hk⟩ := Nat.le.dest hij + revert i j + induction k with + | zero => + rintro i j hij hj hik + obtain rfl : i = j := by omega + have : mkOfLe ⟨i, Nat.lt_add_one_of_le hj⟩ ⟨i, Nat.lt_add_one_of_le hj⟩ (by omega) = + [1].const [0] 0 ≫ [0].const [n] ⟨i, Nat.lt_add_one_of_le hj⟩ := Hom.ext_one_left _ _ + rw [this] + let α : (strArrowMk₂ ([0].const [n] ⟨i, Nat.lt_add_one_of_le hj⟩) (by omega)) ⟶ + (strArrowMk₂ ([1].const [0] 0 ≫ [0].const [n] ⟨i, Nat.lt_add_one_of_le hj⟩) (by omega)) := + StructuredArrow.homMk (([1].const [0] 0).op) (by simp; rfl) + have nat := congr_fun (s.π.naturality α) x + dsimp only [Fin.val_zero, Nat.add_zero, id_eq, Int.reduceNeg, Int.Nat.cast_ofNat_Int, + Int.reduceAdd, Fin.eta, comp_obj, StructuredArrow.proj_obj, op_obj, const_obj_obj, + const_obj_map, types_comp_apply, types_id_apply, Functor.comp_map, StructuredArrow.proj_map, + op_map] at nat + rw [nat, op_comp, Functor.map_comp] + simp only [types_comp_apply] + refine congrArg (X.map ([1].const [0] 0).op) ?_ + unfold strArrowMk₂ + rw [lift, StrictSegal.spineToSimplex_vertex] + congr + | succ k hk => + intro i j hij hj hik + let α := strArrowMk₂ (mkOfLeComp (n := n) ⟨i, by omega⟩ ⟨i + k, by omega⟩ + ⟨j, by omega⟩ (by simp) + (by simp only [Fin.mk_le_mk]; omega)) (by rfl) + let α₀ := strArrowMk₂ (mkOfLe (n := n) ⟨i + k, by omega⟩ ⟨j, by omega⟩ + (by simp only [Fin.mk_le_mk]; omega)) (by simp) + let α₁ := strArrowMk₂ (mkOfLe (n := n) ⟨i, by omega⟩ ⟨j, by omega⟩ + (by simp only [Fin.mk_le_mk]; omega)) (by simp) + let α₂ := strArrowMk₂ (mkOfLe (n := n) ⟨i, by omega⟩ ⟨i + k, by omega⟩ (by simp)) (by simp) + let β₀ : α ⟶ α₀ := StructuredArrow.homMk ((mkOfSucc 1).op) (Quiver.Hom.unop_inj + (by ext x; fin_cases x <;> rfl)) + let β₁ : α ⟶ α₁ := StructuredArrow.homMk ((δ 1).op) (Quiver.Hom.unop_inj + (by ext x; fin_cases x <;> rfl)) + let β₂ : α ⟶ α₂ := StructuredArrow.homMk ((mkOfSucc 0).op) (Quiver.Hom.unop_inj + (by ext x; fin_cases x <;> rfl)) + have h₀ : X.map α₀.hom (lift s x) = s.π.app α₀ x := by + obtain rfl : j = (i + k) + 1 := by omega + exact fac_aux₁ _ _ _ _ (by omega) + have h₂ : X.map α₂.hom (lift s x) = s.π.app α₂ x := + hk i (i + k) (by simp) (by omega) rfl + change X.map α₁.hom (lift s x) = s.π.app α₁ x + have : X.map α.hom (lift s x) = s.π.app α x := by + apply StrictSegal.spineInjective + apply Path.ext' + intro t + dsimp only [spineEquiv] + rw [Equiv.coe_fn_mk, spine_arrow, spine_arrow, + ← FunctorToTypes.map_comp_apply] + match t with + | 0 => + have : α.hom ≫ (mkOfSucc 0).op = α₂.hom := + Quiver.Hom.unop_inj (by ext x ; fin_cases x <;> rfl) + rw [this, h₂, ← congr_fun (s.w β₂) x] + rfl + | 1 => + have : α.hom ≫ (mkOfSucc 1).op = α₀.hom := + Quiver.Hom.unop_inj (by ext x ; fin_cases x <;> rfl) + rw [this, h₀, ← congr_fun (s.w β₀) x] + rfl + rw [← StructuredArrow.w β₁, FunctorToTypes.map_comp_apply, this, ← s.w β₁] + dsimp + +lemma fac_aux₃ {n : ℕ} + (s : Cone (proj (op [n]) (Truncated.inclusion 2).op ⋙ (Truncated.inclusion 2).op ⋙ X)) + (x : s.pt) (φ : [1] ⟶ [n]) : + X.map φ.op (lift s x) = s.π.app (strArrowMk₂ φ (by omega)) x := by + obtain ⟨i, j, hij, rfl⟩ : ∃ i j hij, φ = mkOfLe i j hij := + ⟨φ.toOrderHom 0, φ.toOrderHom 1, φ.toOrderHom.monotone (by simp), + Hom.ext_one_left _ _ rfl rfl⟩ + exact fac_aux₂ _ _ _ _ _ _ (by omega) + +end isPointwiseRightKanExtensionAt + +open Truncated + +open isPointwiseRightKanExtensionAt in +/-- A strict Segal simplicial set is 2-coskeletal. -/ +noncomputable def isPointwiseRightKanExtensionAt (n : ℕ) : + (rightExtensionInclusion X 2).IsPointwiseRightKanExtensionAt ⟨[n]⟩ where + lift s x := lift (X := X) s x + fac s j := by + ext x + obtain ⟨⟨i, hi⟩, ⟨f : _ ⟶ _⟩, rfl⟩ := j.mk_surjective + obtain ⟨i, rfl⟩ : ∃ j, SimplexCategory.mk j = i := ⟨_, i.mk_len⟩ + dsimp at hi ⊢ + apply StrictSegal.spineInjective + dsimp + ext k + · dsimp only [spineEquiv, Equiv.coe_fn_mk] + erw [spine_map_vertex] + rw [spine_spineToSimplex, spine_vertex] + let α : strArrowMk₂ f hi ⟶ strArrowMk₂ ([0].const [n] (f.toOrderHom k)) (by omega) := + StructuredArrow.homMk (([0].const _ (by exact k)).op) (by simp; rfl) + exact congr_fun (s.w α).symm x + · dsimp only [spineEquiv, Equiv.coe_fn_mk, spine_arrow] + rw [← FunctorToTypes.map_comp_apply] + let α : strArrowMk₂ f hi ⟶ strArrowMk₂ (mkOfSucc k ≫ f) (by omega) := + StructuredArrow.homMk (mkOfSucc k).op (by simp; rfl) + exact (isPointwiseRightKanExtensionAt.fac_aux₃ _ _ _ _).trans (congr_fun (s.w α).symm x) + uniq s m hm := by + ext x + apply StrictSegal.spineInjective (X := X) + dsimp [spineEquiv] + rw [StrictSegal.spine_spineToSimplex] + ext i + · exact congr_fun (hm (StructuredArrow.mk (Y := op [0]₂) ([0].const [n] i).op)) x + · exact congr_fun (hm (.mk (Y := op [1]₂) (.op (mkOfLe _ _ (Fin.castSucc_le_succ i))))) x + +/-- Since `StrictSegal.isPointwiseRightKanExtensionAt` proves that the appropriate +cones are limit cones, `rightExtensionInclusion X 2` is a pointwise right Kan extension.-/ +noncomputable def isPointwiseRightKanExtension : + (rightExtensionInclusion X 2).IsPointwiseRightKanExtension := + fun Δ => isPointwiseRightKanExtensionAt X Δ.unop.len + +theorem isRightKanExtension : + X.IsRightKanExtension (𝟙 ((inclusion 2).op ⋙ X)) := + RightExtension.IsPointwiseRightKanExtension.isRightKanExtension + (isPointwiseRightKanExtension X) + +/-- When `X` is `StrictSegal`, `X` is 2-coskeletal. -/ +instance isCoskeletal : SimplicialObject.IsCoskeletal X 2 where + isRightKanExtension := isRightKanExtension X + +end StrictSegal + +end + +end SSet + +namespace CategoryTheory + +namespace Nerve + +open SSet + +example (C : Type u) [Category.{v} C] : + SimplicialObject.IsCoskeletal (nerve C) 2 := by infer_instance + +/-- The essential data of the nerve functor is contained in the 2-truncation, which is +recorded by the composite functor `nerveFunctor₂`.-/ +def nerveFunctor₂ : Cat.{v, u} ⥤ SSet.Truncated 2 := nerveFunctor ⋙ truncation 2 + +/-- The natural isomorphism between `nerveFunctor` and `nerveFunctor₂ ⋙ Truncated.cosk 2` whose +components `nerve C ≅ (Truncated.cosk 2).obj (nerveFunctor₂.obj C)` shows that nerves of categories +are 2-coskeletal.-/ +noncomputable def cosk₂Iso : nerveFunctor.{v, u} ≅ nerveFunctor₂.{v, u} ⋙ Truncated.cosk 2 := + NatIso.ofComponents (fun C ↦ (nerve C).isoCoskOfIsCoskeletal 2) + (fun _ ↦ (coskAdj 2).unit.naturality _) + +end Nerve + +end CategoryTheory diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Nerve.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Nerve.lean index c82254637aaf3..a24814b5a2a14 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Nerve.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Nerve.lean @@ -37,7 +37,7 @@ instance {C : Type*} [Category C] {Δ : SimplexCategoryᵒᵖ} : Category ((nerv /-- The nerve of a category, as a functor `Cat ⥤ SSet` -/ @[simps] -def nerveFunctor : Cat ⥤ SSet where +def nerveFunctor : Cat.{v, u} ⥤ SSet where obj C := nerve C map F := { app := fun _ => (F.mapComposableArrows _).obj } diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean index 33e4d07da1bce..ff920be3bfa29 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean @@ -64,6 +64,13 @@ def spine (n : ℕ) (Δ : X _[n]) : X.Path n where simp only [← FunctorToTypes.map_comp_apply, ← op_comp] rw [SimplexCategory.δ_zero_mkOfSucc] +lemma spine_map_vertex {n : ℕ} (x : X _[n]) {m : ℕ} (φ : ([m] : SimplexCategory) ⟶ [n]) + (i : Fin (m + 1)) : + (spine X m (X.map φ.op x)).vertex i = (spine X n x).vertex (φ.toOrderHom i) := by + dsimp [spine] + rw [← FunctorToTypes.map_comp_apply] + rfl + lemma spine_map_subinterval {n : ℕ} (j l : ℕ) (hjl : j + l ≤ n) (Δ : X _[n]) : X.spine l (X.map (subinterval j l (by omega)).op Δ) = (X.spine n Δ).interval j l (by omega) := by diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean b/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean index 7436832a79255..d45282b986c0d 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean @@ -19,6 +19,10 @@ Examples of `StrictSegal` simplicial sets are given by nerves of categories. TODO: Show that these are the only examples: that a `StrictSegal` simplicial set is isomorphic to the nerve of its homotopy category. + +`StrictSegal` simplicial sets have an important property of being 2-coskeletal which is proven +in `Mathlib.AlgebraicTopology.SimplicialSet.Coskeletal`. + -/ universe v u From 7a05b449ac13af0118fd6b2c133d08e33e7e74ad Mon Sep 17 00:00:00 2001 From: FMLJohn <417963616@qq.com> Date: Wed, 4 Dec 2024 20:23:40 +0000 Subject: [PATCH 514/829] feat(RingTheory/Polynomial/Hilbert): `Polynomial.exists_unique_hilbertPoly` and `Polynomial.hilbertPoly_mul_one_sub_pow_add` (#19404) Co-authored-by: Li --- .../RingTheory/Polynomial/HilbertPoly.lean | 54 +++++++++++++++++-- Mathlib/RingTheory/PowerSeries/WellKnown.lean | 12 +++++ 2 files changed, 62 insertions(+), 4 deletions(-) diff --git a/Mathlib/RingTheory/Polynomial/HilbertPoly.lean b/Mathlib/RingTheory/Polynomial/HilbertPoly.lean index 302a9660f7465..7aa04415b3b93 100644 --- a/Mathlib/RingTheory/Polynomial/HilbertPoly.lean +++ b/Mathlib/RingTheory/Polynomial/HilbertPoly.lean @@ -5,6 +5,8 @@ Authors: Fangming Li, Jujian Zhang -/ import Mathlib.Algebra.Polynomial.AlgebraMap import Mathlib.Algebra.Polynomial.Eval.SMul +import Mathlib.Algebra.Polynomial.Roots +import Mathlib.Order.Interval.Set.Infinite import Mathlib.RingTheory.Polynomial.Pochhammer import Mathlib.RingTheory.PowerSeries.WellKnown import Mathlib.Tactic.FieldSimp @@ -17,9 +19,9 @@ given any `p : F[X]` and `d : ℕ`, there exists some `h : F[X]` such that for a `n : ℕ`, `h(n)` is equal to the coefficient of `Xⁿ` in the power series expansion of `p/(1 - X)ᵈ`. This `h` is unique and is denoted as `Polynomial.hilbertPoly p d`. -For example, given `d : ℕ`, the power series expansion of `1/(1-X)ᵈ⁺¹` in `F[X]` is -`Σₙ ((d + n).choose d)Xⁿ`, which equals `Σₙ ((n + 1)···(n + d)/d!)Xⁿ` and hence -`Polynomial.hilbertPoly (1 : F[X]) d` is the polynomial `(n + 1)···(n + d)/d!`. Note that +For example, given `d : ℕ`, the power series expansion of `1/(1-X)ᵈ⁺¹` in `F[X]` +is `Σₙ ((d + n).choose d)Xⁿ`, which equals `Σₙ ((n + 1)···(n + d)/d!)Xⁿ` and hence +`Polynomial.hilbertPoly (1 : F[X]) (d + 1)` is the polynomial `(n + 1)···(n + d)/d!`. Note that if `d! = 0` in `F`, then the polynomial `(n + 1)···(n + d)/d!` no longer works, so we do not want the characteristic of `F` to be divisible by `d!`. As `Polynomial.hilbertPoly` may take any `p : F[X]` and `d : ℕ` as its inputs, it is necessary for us to assume that `CharZero F`. @@ -96,13 +98,15 @@ lemma hilbertPoly_X_pow_succ (d k : ℕ) : hilbertPoly ((X : F[X]) ^ k) (d + 1) = preHilbertPoly F d k := by delta hilbertPoly; simp +variable [CharZero F] + /-- The key property of Hilbert polynomials. If `F` is a field with characteristic `0`, `p : F[X]` and `d : ℕ`, then for any large enough `n : ℕ`, `(Polynomial.hilbertPoly p d).eval (n : F)` equals the coefficient of `Xⁿ` in the power series expansion of `p/(1 - X)ᵈ`. -/ theorem coeff_mul_invOneSubPow_eq_hilbertPoly_eval - [CharZero F] {p : F[X]} (d : ℕ) {n : ℕ} (hn : p.natDegree < n) : + {p : F[X]} (d : ℕ) {n : ℕ} (hn : p.natDegree < n) : PowerSeries.coeff F n (p * (invOneSubPow F d)) = (hilbertPoly p d).eval (n : F) := by delta hilbertPoly; induction d with | zero => simp only [invOneSubPow_zero, Units.val_one, mul_one, coeff_coe, eval_zero] @@ -127,4 +131,46 @@ theorem coeff_mul_invOneSubPow_eq_hilbertPoly_eval rw [hx.2, zero_mul] · rw [add_comm, Nat.add_sub_assoc (h_le ⟨x, hx⟩), succ_eq_add_one, add_tsub_cancel_right] +/-- +The polynomial satisfying the key property of `Polynomial.hilbertPoly p d` is unique. In other +words, if `h : F[X]` and there exists some `N : ℕ` such that for any number `n : ℕ` bigger than +`N` we have `PowerSeries.coeff F n (p * (invOneSubPow F d)) = h.eval (n : F)`, then `h` is exactly +`Polynomial.hilbertPoly p d`. +-/ +theorem exists_unique_hilbertPoly (p : F[X]) (d : ℕ) : + ∃! (h : F[X]), (∃ (N : ℕ), (∀ (n : ℕ) (_ : N < n), + PowerSeries.coeff F n (p * (invOneSubPow F d)) = h.eval (n : F))) := by + use hilbertPoly p d; constructor + · use p.natDegree + exact fun n => coeff_mul_invOneSubPow_eq_hilbertPoly_eval d + · rintro h ⟨N, hhN⟩ + apply eq_of_infinite_eval_eq h (hilbertPoly p d) + apply ((Set.Ioi_infinite (max N p.natDegree)).image cast_injective.injOn).mono + rintro x ⟨n, hn, rfl⟩ + simp only [Set.mem_Ioi, sup_lt_iff, Set.mem_setOf_eq] at hn ⊢ + rw [← coeff_mul_invOneSubPow_eq_hilbertPoly_eval d hn.2, hhN n hn.1] + +lemma hilbertPoly_mul_one_sub_succ (p : F[X]) (d : ℕ) : + hilbertPoly (p * (1 - X)) (d + 1) = hilbertPoly p d := by + apply eq_of_infinite_eval_eq + apply ((Set.Ioi_infinite (p * (1 - X)).natDegree).image cast_injective.injOn).mono + rintro x ⟨n, hn, rfl⟩ + simp only [Set.mem_setOf_eq] + by_cases hp : p = 0 + · simp only [hp, zero_mul, hilbertPoly_zero_nat] + · simp only [Set.mem_Ioi] at hn + have hpn : p.natDegree < n := by + suffices (1 : F[X]) - X ≠ 0 from lt_of_add_right_lt (natDegree_mul hp this ▸ hn) + rw [sub_ne_zero, ne_eq, ext_iff, not_forall] + use 0 + simp + simp only [hn, ← coeff_mul_invOneSubPow_eq_hilbertPoly_eval, coe_mul, coe_sub, coe_one, coe_X, + mul_assoc, hpn, ← one_sub_pow_mul_invOneSubPow_val_add_eq_invOneSubPow_val F d 1, pow_one] + +lemma hilbertPoly_mul_one_sub_pow_add (p : F[X]) (d e : ℕ) : + hilbertPoly (p * (1 - X) ^ e) (d + e) = hilbertPoly p d := by + induction e with + | zero => simp + | succ e he => rw [pow_add, pow_one, ← mul_assoc, ← add_assoc, hilbertPoly_mul_one_sub_succ, he] + end Polynomial diff --git a/Mathlib/RingTheory/PowerSeries/WellKnown.lean b/Mathlib/RingTheory/PowerSeries/WellKnown.lean index ca5bd0b442330..783e6348c6aa9 100644 --- a/Mathlib/RingTheory/PowerSeries/WellKnown.lean +++ b/Mathlib/RingTheory/PowerSeries/WellKnown.lean @@ -159,6 +159,18 @@ theorem mk_add_choose_mul_one_sub_pow_eq_one : (mk fun n ↦ Nat.choose (d + n) d : S⟦X⟧) * ((1 - X) ^ (d + 1)) = 1 := (invOneSubPow S (d + 1)).val_inv +theorem invOneSubPow_add (e : ℕ) : + invOneSubPow S (d + e) = invOneSubPow S d * invOneSubPow S e := by + simp_rw [invOneSubPow_eq_inv_one_sub_pow, pow_add] + +theorem one_sub_pow_mul_invOneSubPow_val_add_eq_invOneSubPow_val (e : ℕ) : + (1 - X) ^ e * (invOneSubPow S (d + e)).val = (invOneSubPow S d).val := by + simp [invOneSubPow_add, Units.val_mul, mul_comm, mul_assoc, ← invOneSubPow_inv_eq_one_sub_pow] + +theorem one_sub_pow_add_mul_invOneSubPow_val_eq_one_sub_pow (e : ℕ) : + (1 - X) ^ (d + e) * (invOneSubPow S e).val = (1 - X) ^ d := by + simp [pow_add, mul_assoc, ← invOneSubPow_inv_eq_one_sub_pow S e] + end invOneSubPow section Field From 9bec372f2c4bd188669f930fa1b74bd13a52131b Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Wed, 4 Dec 2024 21:50:19 +0000 Subject: [PATCH 515/829] chore: update Mathlib dependencies 2024-12-04 (#19732) 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 5edf45f8e3379..9545da43cc321 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "7805acf1864ba1a2e359f86a8f092ccf1438ad83", + "rev": "c016aa9938c4cedc9b7066099f99bcae1b1af625", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From 3961c23f187da92b1208cd35d277443287bb771e Mon Sep 17 00:00:00 2001 From: Mac Malone Date: Wed, 4 Dec 2024 23:17:54 +0000 Subject: [PATCH 516/829] fix(Cache): always invoke `lake` to fetch ProofWidgets (#19728) Always invoke `lake` to fetch the latest ProofWidgets cloud release on a `lake exe cache get`. This requires `lake` because cache has no simple heuristic to determine whether the ProofWidgets JS is out-of-date. --- Cache/Requests.lean | 23 ++++++++++++++++++----- 1 file changed, 18 insertions(+), 5 deletions(-) diff --git a/Cache/Requests.lean b/Cache/Requests.lean index 295d65cd96693..38f3db11e9406 100644 --- a/Cache/Requests.lean +++ b/Cache/Requests.lean @@ -161,14 +161,27 @@ into the `lean-toolchain` file at the root directory of your project" /-- Fetches the ProofWidgets cloud release and prunes non-JS files. -/ def getProofWidgets (buildDir : FilePath) : IO Unit := do - if (← buildDir.pathExists) then return - -- Unpack the ProofWidgets cloud release (for its `.js` files) + if (← buildDir.pathExists) then + -- Check if the ProofWidgets build is out-of-date via `lake`. + -- This is done through Lake as cache has no simple heuristic + -- to determine whether the ProofWidgets JS is out-of-date. + let exitCode ← (← IO.Process.spawn {cmd := "lake", args := #["-q", "build", "--no-build", "proofwidgets:release"]}).wait + if exitCode == 0 then -- up-to-date + return + else if exitCode == 3 then -- needs fetch (`--no-build` triggered) + pure () + else + throw <| IO.userError s!"Failed to validate ProofWidgets cloud release: lake failed with error code {exitCode}" + -- Download and unpack the ProofWidgets cloud release (for its `.js` files) let exitCode ← (← IO.Process.spawn {cmd := "lake", args := #["-q", "build", "proofwidgets:release"]}).wait if exitCode != 0 then throw <| IO.userError s!"Failed to fetch ProofWidgets cloud release: lake failed with error code {exitCode}" - -- prune non-js ProofWidgets files (e.g., `olean`, `.c`) - IO.FS.removeDirAll (buildDir / "lib") - IO.FS.removeDirAll (buildDir / "ir") + -- Prune non-JS ProofWidgets files (e.g., `olean`, `.c`) + try + IO.FS.removeDirAll (buildDir / "lib") + IO.FS.removeDirAll (buildDir / "ir") + catch e => + throw <| IO.userError s!"Failed to prune ProofWidgets cloud release: {e}" /-- Downloads missing files, and unpacks files. -/ def getFiles (hashMap : IO.HashMap) (forceDownload forceUnpack parallel decompress : Bool) : From 9b2bb8d181d2263d589e05723b81426f235f7e06 Mon Sep 17 00:00:00 2001 From: Mitchell Lee Date: Thu, 5 Dec 2024 00:45:04 +0000 Subject: [PATCH 517/829] doc(Order/BoundedOrder/Lattice): fix typos (#19719) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Change the sentence "Typical examples include `Prop` and `Det α`." to "Typical examples include `Prop` and `Set α`." Also, fix a few other typos. --- Mathlib/Order/BoundedOrder/Lattice.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/Order/BoundedOrder/Lattice.lean b/Mathlib/Order/BoundedOrder/Lattice.lean index 42c1cf7c0ac57..3c06e09a51844 100644 --- a/Mathlib/Order/BoundedOrder/Lattice.lean +++ b/Mathlib/Order/BoundedOrder/Lattice.lean @@ -7,19 +7,19 @@ import Mathlib.Order.BoundedOrder.Basic import Mathlib.Order.Lattice /-! -# bounded lattices +# Bounded lattices This file defines top and bottom elements (greatest and least elements) of a type, the bounded -variants of different kinds of lattices, sets up the typeclass hierarchy between them and provides +variants of different kinds of lattices, sets up the typeclass hierarchy between them, and provides instances for `Prop` and `fun`. ## Common lattices -* Distributive lattices with a bottom element. Notated by `[DistribLattice α] [OrderBot α]` +* Distributive lattices with a bottom element. Notated by `[DistribLattice α] [OrderBot α]`. It captures the properties of `Disjoint` that are common to `GeneralizedBooleanAlgebra` and `DistribLattice` when `OrderBot`. * Bounded and distributive lattice. Notated by `[DistribLattice α] [BoundedOrder α]`. - Typical examples include `Prop` and `Det α`. + Typical examples include `Prop` and `Set α`. -/ open Function OrderDual From 639c1cb9ae7fa522a739834eb1c3af94908ae0c4 Mon Sep 17 00:00:00 2001 From: Paul Lezeau Date: Thu, 5 Dec 2024 01:28:06 +0000 Subject: [PATCH 518/829] feat(Order/Bounds/Basic): Add basic order lemmas (#17595) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds some simple lemmas about `IsLeast/IsGreatest` for images of sets along `StrictMono/StrictAnti` maps. Co-authored-by: Yaël Dillies --- Mathlib/Order/Bounds/Image.lean | 39 +++++++++++++++++++++++++++++++++ 1 file changed, 39 insertions(+) diff --git a/Mathlib/Order/Bounds/Image.lean b/Mathlib/Order/Bounds/Image.lean index 910c466332520..cb97202b9d3ce 100644 --- a/Mathlib/Order/Bounds/Image.lean +++ b/Mathlib/Order/Bounds/Image.lean @@ -196,6 +196,45 @@ theorem map_isLeast : IsLeast s a → IsGreatest (f '' s) (f a) := end Antitone +section StrictMono + +variable [LinearOrder α] [Preorder β] {f : α → β} {a : α} {s : Set α} + +lemma StrictMono.mem_upperBounds_image (hf : StrictMono f) : + f a ∈ upperBounds (f '' s) ↔ a ∈ upperBounds s := by simp [upperBounds, hf.le_iff_le] + +lemma StrictMono.mem_lowerBounds_image (hf : StrictMono f) : + f a ∈ lowerBounds (f '' s) ↔ a ∈ lowerBounds s := by simp [lowerBounds, hf.le_iff_le] + +lemma StrictMono.map_isLeast (hf : StrictMono f) : IsLeast (f '' s) (f a) ↔ IsLeast s a := by + simp [IsLeast, hf.injective.eq_iff, hf.mem_lowerBounds_image] + +lemma StrictMono.map_isGreatest (hf : StrictMono f) : + IsGreatest (f '' s) (f a) ↔ IsGreatest s a := by + simp [IsGreatest, hf.injective.eq_iff, hf.mem_upperBounds_image] + +end StrictMono + +section StrictAnti + +variable [LinearOrder α] [Preorder β] {f : α → β} {a : α} {s : Set α} + +lemma StrictAnti.mem_upperBounds_image (hf : StrictAnti f) : + f a ∈ upperBounds (f '' s) ↔ a ∈ lowerBounds s := by + simp [upperBounds, lowerBounds, hf.le_iff_le] + +lemma StrictAnti.mem_lowerBounds_image (hf : StrictAnti f) : + f a ∈ lowerBounds (f '' s) ↔ a ∈ upperBounds s := by + simp [upperBounds, lowerBounds, hf.le_iff_le] + +lemma StrictAnti.map_isLeast (hf : StrictAnti f) : IsLeast (f '' s) (f a) ↔ IsGreatest s a := by + simp [IsLeast, IsGreatest, hf.injective.eq_iff, hf.mem_lowerBounds_image] + +lemma StrictAnti.map_isGreatest (hf : StrictAnti f) : IsGreatest (f '' s) (f a) ↔ IsLeast s a := by + simp [IsLeast, IsGreatest, hf.injective.eq_iff, hf.mem_upperBounds_image] + +end StrictAnti + section Image2 variable [Preorder α] [Preorder β] [Preorder γ] {f : α → β → γ} {s : Set α} {t : Set β} {a : α} From 80e107a996a9adfbf1aa1f555e99dfe9137f5dcd Mon Sep 17 00:00:00 2001 From: grunweg Date: Thu, 5 Dec 2024 01:28:08 +0000 Subject: [PATCH 519/829] chore: add more gcongr attributes (#17610) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR tags (analogues of) lemmas tagged with `mono` also with `gcongr`. Co-authored-by: Yaël Dillies --- Mathlib/Algebra/Group/Subgroup/Basic.lean | 1 + Mathlib/Algebra/Group/Submonoid/Pointwise.lean | 11 +++-------- Mathlib/Algebra/Lie/Solvable.lean | 2 +- Mathlib/Algebra/Lie/Submodule.lean | 2 +- Mathlib/Algebra/Module/Submodule/Bilinear.lean | 2 +- Mathlib/Algebra/Ring/Subring/Basic.lean | 15 ++++++++++++++- Mathlib/Algebra/Ring/Subsemiring/Basic.lean | 2 +- Mathlib/Analysis/BoxIntegral/Box/Basic.lean | 2 +- .../Analysis/BoxIntegral/Partition/Filter.lean | 6 +++--- .../FunctionalSpaces/SobolevInequality.lean | 5 ++++- .../SpecialFunctions/Trigonometric/Arctan.lean | 7 +++++++ Mathlib/Combinatorics/SimpleGraph/Basic.lean | 2 +- Mathlib/Combinatorics/SimpleGraph/Clique.lean | 4 ++-- Mathlib/Combinatorics/SimpleGraph/Subgraph.lean | 8 ++++---- Mathlib/Deprecated/Subgroup.lean | 1 + .../LinearAlgebra/AffineSpace/AffineSubspace.lean | 2 +- Mathlib/LinearAlgebra/Basis/Flag.lean | 8 +++++++- .../LinearAlgebra/Projectivization/Subspace.lean | 3 +++ .../MeasureTheory/Function/LpSeminorm/Basic.lean | 6 +++--- Mathlib/MeasureTheory/Function/SimpleFunc.lean | 2 +- Mathlib/MeasureTheory/Integral/Bochner.lean | 2 +- .../Integral/DominatedConvergence.lean | 2 +- Mathlib/MeasureTheory/Integral/Lebesgue.lean | 2 +- Mathlib/Topology/NhdsSet.lean | 2 +- 24 files changed, 64 insertions(+), 35 deletions(-) diff --git a/Mathlib/Algebra/Group/Subgroup/Basic.lean b/Mathlib/Algebra/Group/Subgroup/Basic.lean index c6b11e2b701e3..f828e651230b1 100644 --- a/Mathlib/Algebra/Group/Subgroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subgroup/Basic.lean @@ -479,6 +479,7 @@ theorem normalClosure_le_normal {N : Subgroup G} [N.Normal] (h : s ⊆ N) : norm theorem normalClosure_subset_iff {N : Subgroup G} [N.Normal] : s ⊆ N ↔ normalClosure s ≤ N := ⟨normalClosure_le_normal, Set.Subset.trans subset_normalClosure⟩ +@[gcongr] theorem normalClosure_mono {s t : Set G} (h : s ⊆ t) : normalClosure s ≤ normalClosure t := normalClosure_le_normal (Set.Subset.trans h subset_normalClosure) diff --git a/Mathlib/Algebra/Group/Submonoid/Pointwise.lean b/Mathlib/Algebra/Group/Submonoid/Pointwise.lean index 596c1f0f1ca02..567936d4f37ef 100644 --- a/Mathlib/Algebra/Group/Submonoid/Pointwise.lean +++ b/Mathlib/Algebra/Group/Submonoid/Pointwise.lean @@ -587,15 +587,10 @@ theorem bot_mul (S : AddSubmonoid R) : ⊥ * S = ⊥ := variable {M N P Q : AddSubmonoid R} -@[mono] -theorem mul_le_mul (hmp : M ≤ P) (hnq : N ≤ Q) : M * N ≤ P * Q := - smul_le_smul hmp hnq +@[mono, gcongr] lemma mul_le_mul (hmp : M ≤ P) (hnq : N ≤ Q) : M * N ≤ P * Q := smul_le_smul hmp hnq -theorem mul_le_mul_left (h : M ≤ N) : M * P ≤ N * P := - smul_le_smul_left h - -theorem mul_le_mul_right (h : N ≤ P) : M * N ≤ M * P := - smul_le_smul_right h +@[gcongr] lemma mul_le_mul_left (h : M ≤ N) : M * P ≤ N * P := smul_le_smul_left h +@[gcongr] lemma mul_le_mul_right (h : N ≤ P) : M * N ≤ M * P := smul_le_smul_right h theorem mul_subset_mul : (↑M : Set R) * (↑N : Set R) ⊆ (↑(M * N) : Set R) := smul_subset_smul diff --git a/Mathlib/Algebra/Lie/Solvable.lean b/Mathlib/Algebra/Lie/Solvable.lean index 1ce5fd8fdcf58..4a9632102afc4 100644 --- a/Mathlib/Algebra/Lie/Solvable.lean +++ b/Mathlib/Algebra/Lie/Solvable.lean @@ -77,7 +77,7 @@ theorem derivedSeriesOfIdeal_add (k l : ℕ) : D (k + l) I = D k (D l I) := by | zero => rw [Nat.zero_add, derivedSeriesOfIdeal_zero] | succ k ih => rw [Nat.succ_add k l, derivedSeriesOfIdeal_succ, derivedSeriesOfIdeal_succ, ih] -@[mono] +@[gcongr, mono] theorem derivedSeriesOfIdeal_le {I J : LieIdeal R L} {k l : ℕ} (h₁ : I ≤ J) (h₂ : l ≤ k) : D k I ≤ D l J := by revert l; induction' k with k ih <;> intro l h₂ diff --git a/Mathlib/Algebra/Lie/Submodule.lean b/Mathlib/Algebra/Lie/Submodule.lean index 1201232787862..32a16eb6c461c 100644 --- a/Mathlib/Algebra/Lie/Submodule.lean +++ b/Mathlib/Algebra/Lie/Submodule.lean @@ -838,7 +838,7 @@ theorem comap_incl_eq_bot : N₂.comap N.incl = ⊥ ↔ N ⊓ N₂ = ⊥ := by inf_coe_toSubmodule] rw [← Submodule.disjoint_iff_comap_eq_bot, disjoint_iff] -@[mono] +@[gcongr, mono] theorem map_mono (h : N ≤ N₂) : N.map f ≤ N₂.map f := Set.image_subset _ h diff --git a/Mathlib/Algebra/Module/Submodule/Bilinear.lean b/Mathlib/Algebra/Module/Submodule/Bilinear.lean index 1564df8114cd7..0f4c48b6a1522 100644 --- a/Mathlib/Algebra/Module/Submodule/Bilinear.lean +++ b/Mathlib/Algebra/Module/Submodule/Bilinear.lean @@ -84,7 +84,7 @@ theorem map₂_bot_left (f : M →ₗ[R] N →ₗ[R] P) (q : Submodule R N) : ma rw [Submodule.mem_bot] at hm ⊢ rw [hm, LinearMap.map_zero₂] -@[mono] +@[gcongr, mono] theorem map₂_le_map₂ {f : M →ₗ[R] N →ₗ[R] P} {p₁ p₂ : Submodule R M} {q₁ q₂ : Submodule R N} (hp : p₁ ≤ p₂) (hq : q₁ ≤ q₂) : map₂ f p₁ q₁ ≤ map₂ f p₂ q₂ := map₂_le.2 fun _m hm _n hn => apply_mem_map₂ _ (hp hm) (hq hn) diff --git a/Mathlib/Algebra/Ring/Subring/Basic.lean b/Mathlib/Algebra/Ring/Subring/Basic.lean index 26820e434059b..d81eb33fca30f 100644 --- a/Mathlib/Algebra/Ring/Subring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subring/Basic.lean @@ -66,6 +66,7 @@ variable {R : Type u} {S : Type v} {T : Type w} [Ring R] variable [Ring S] [Ring T] namespace Subring +variable {s t : Subring R} -- Porting note: there is no `Subring.toSubmonoid` but we can't define it because there is a -- projection `s.toSubmonoid` @@ -78,6 +79,12 @@ theorem toSubsemiring_strictMono : StrictMono (toSubsemiring : Subring R → Sub theorem toSubsemiring_mono : Monotone (toSubsemiring : Subring R → Subsemiring R) := toSubsemiring_strictMono.monotone +@[gcongr] +lemma toSubsemiring_lt_toSubsemiring (hst : s < t) : s.toSubsemiring < t.toSubsemiring := hst + +@[gcongr] +lemma toSubsemiring_le_toSubsemiring (hst : s ≤ t) : s.toSubsemiring ≤ t.toSubsemiring := hst + @[mono] theorem toAddSubgroup_strictMono : StrictMono (toAddSubgroup : Subring R → AddSubgroup R) := fun _ _ => id @@ -86,6 +93,12 @@ theorem toAddSubgroup_strictMono : StrictMono (toAddSubgroup : Subring R → Add theorem toAddSubgroup_mono : Monotone (toAddSubgroup : Subring R → AddSubgroup R) := toAddSubgroup_strictMono.monotone +@[gcongr] +lemma toAddSubgroup_lt_toAddSubgroup (hst : s < t) : s.toAddSubgroup < t.toAddSubgroup := hst + +@[gcongr] +lemma toAddSubgroup_le_toAddSubgroup (hst : s ≤ t) : s.toAddSubgroup ≤ t.toAddSubgroup := hst + @[mono] theorem toSubmonoid_strictMono : StrictMono (fun s : Subring R => s.toSubmonoid) := fun _ _ => id @@ -656,7 +669,7 @@ theorem coe_prod (s : Subring R) (t : Subring S) : theorem mem_prod {s : Subring R} {t : Subring S} {p : R × S} : p ∈ s.prod t ↔ p.1 ∈ s ∧ p.2 ∈ t := Iff.rfl -@[mono] +@[gcongr, mono] theorem prod_mono ⦃s₁ s₂ : Subring R⦄ (hs : s₁ ≤ s₂) ⦃t₁ t₂ : Subring S⦄ (ht : t₁ ≤ t₂) : s₁.prod t₁ ≤ s₂.prod t₂ := Set.prod_mono hs ht diff --git a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean index 65c4d305a67a1..88035633f260c 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean @@ -589,7 +589,7 @@ theorem mem_prod {s : Subsemiring R} {t : Subsemiring S} {p : R × S} : p ∈ s.prod t ↔ p.1 ∈ s ∧ p.2 ∈ t := Iff.rfl -@[mono] +@[gcongr, mono] theorem prod_mono ⦃s₁ s₂ : Subsemiring R⦄ (hs : s₁ ≤ s₂) ⦃t₁ t₂ : Subsemiring S⦄ (ht : t₁ ≤ t₂) : s₁.prod t₁ ≤ s₂.prod t₂ := Set.prod_mono hs ht diff --git a/Mathlib/Analysis/BoxIntegral/Box/Basic.lean b/Mathlib/Analysis/BoxIntegral/Box/Basic.lean index fa0d86469c73b..622d8b9eb3b5d 100644 --- a/Mathlib/Analysis/BoxIntegral/Box/Basic.lean +++ b/Mathlib/Analysis/BoxIntegral/Box/Basic.lean @@ -357,7 +357,7 @@ def face {n} (I : Box (Fin (n + 1))) (i : Fin (n + 1)) : Box (Fin n) := theorem face_mk {n} (l u : Fin (n + 1) → ℝ) (h : ∀ i, l i < u i) (i : Fin (n + 1)) : face ⟨l, u, h⟩ i = ⟨l ∘ Fin.succAbove i, u ∘ Fin.succAbove i, fun _ ↦ h _⟩ := rfl -@[mono] +@[gcongr, mono] theorem face_mono {n} {I J : Box (Fin (n + 1))} (h : I ≤ J) (i : Fin (n + 1)) : face I i ≤ face J i := fun _ hx _ ↦ Ioc_subset_Ioc ((le_iff_bounds.1 h).1 _) ((le_iff_bounds.1 h).2 _) (hx _) diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean b/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean index c1b1d7ea49571..f9ca9c7272245 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean @@ -408,19 +408,19 @@ nonrec theorem RCond.min {ι : Type*} {r₁ r₂ : (ι → ℝ) → Ioi (0 : ℝ (h₂ : l.RCond r₂) : l.RCond fun x => min (r₁ x) (r₂ x) := fun hR x => congr_arg₂ min (h₁ hR x) (h₂ hR x) -@[mono] +@[gcongr, mono] theorem toFilterDistortion_mono (I : Box ι) (h : l₁ ≤ l₂) (hc : c₁ ≤ c₂) : l₁.toFilterDistortion I c₁ ≤ l₂.toFilterDistortion I c₂ := iInf_mono fun _ => iInf_mono' fun hr => ⟨hr.mono h, principal_mono.2 fun _ => MemBaseSet.mono I h hc fun _ _ => le_rfl⟩ -@[mono] +@[gcongr, mono] theorem toFilter_mono (I : Box ι) {l₁ l₂ : IntegrationParams} (h : l₁ ≤ l₂) : l₁.toFilter I ≤ l₂.toFilter I := iSup_mono fun _ => toFilterDistortion_mono I h le_rfl -@[mono] +@[gcongr, mono] theorem toFilteriUnion_mono (I : Box ι) {l₁ l₂ : IntegrationParams} (h : l₁ ≤ l₂) (π₀ : Prepartition I) : l₁.toFilteriUnion I π₀ ≤ l₂.toFilteriUnion I π₀ := iSup_mono fun _ => inf_le_inf_right _ <| toFilterDistortion_mono _ h le_rfl diff --git a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean index e18e2e98d3690..660294ee97883 100644 --- a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean +++ b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean @@ -336,8 +336,11 @@ theorem lintegral_pow_le_pow_lintegral_fderiv_aux [Fintype ι] · exact hu.comp (by convert contDiff_update 1 x i) · exact h2u.comp_isClosedEmbedding (isClosedEmbedding_update x i) _ ≤ ∫⁻ xᵢ, (‖fderiv ℝ u (update x i xᵢ)‖₊ : ℝ≥0∞) := ?_ - gcongr with y; swap + gcongr · exact Measure.restrict_le_self + intro y + dsimp + gcongr -- bound the derivative which appears calc ‖deriv (u ∘ update x i) y‖₊ = ‖fderiv ℝ u (update x i y) (deriv (update x i) y)‖₊ := by rw [fderiv_comp_deriv _ (hu.differentiable le_rfl).differentiableAt diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Arctan.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Arctan.lean index 6919f9dcbb728..7e4c7f3d67170 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Arctan.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Arctan.lean @@ -145,6 +145,13 @@ theorem arctan_zero : arctan 0 = 0 := by simp [arctan_eq_arcsin] @[mono] theorem arctan_strictMono : StrictMono arctan := tanOrderIso.symm.strictMono +@[gcongr] +lemma arctan_lt_arctan {x y : ℝ} (hxy : x < y) : arctan x < arctan y := arctan_strictMono hxy + +@[gcongr] +lemma arctan_le_arctan {x y : ℝ} (hxy : x ≤ y) : arctan x ≤ arctan y := + arctan_strictMono.monotone hxy + theorem arctan_injective : arctan.Injective := arctan_strictMono.injective @[simp] diff --git a/Mathlib/Combinatorics/SimpleGraph/Basic.lean b/Mathlib/Combinatorics/SimpleGraph/Basic.lean index 1da6a51b3e613..fbf6473d0fddf 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Basic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Basic.lean @@ -603,7 +603,7 @@ theorem fromEdgeSet_sdiff (s t : Set (Sym2 V)) : ext v w constructor <;> simp +contextual -@[mono] +@[gcongr, mono] theorem fromEdgeSet_mono {s t : Set (Sym2 V)} (h : s ⊆ t) : fromEdgeSet s ≤ fromEdgeSet t := by rintro v w simp +contextual only [fromEdgeSet_adj, Ne, not_false_iff, diff --git a/Mathlib/Combinatorics/SimpleGraph/Clique.lean b/Mathlib/Combinatorics/SimpleGraph/Clique.lean index dd588e24a5611..ea2c7d7ec4dbc 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Clique.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Clique.lean @@ -493,7 +493,7 @@ theorem cliqueSet_eq_empty_iff : G.cliqueSet n = ∅ ↔ G.CliqueFree n := by protected alias ⟨_, CliqueFree.cliqueSet⟩ := cliqueSet_eq_empty_iff -@[mono] +@[gcongr, mono] theorem cliqueSet_mono (h : G ≤ H) : G.cliqueSet n ⊆ H.cliqueSet n := fun _ ↦ IsNClique.mono h @@ -636,7 +636,7 @@ theorem card_cliqueFinset_le : #(G.cliqueFinset n) ≤ (card α).choose n := by variable [DecidableRel H.Adj] -@[mono] +@[gcongr, mono] theorem cliqueFinset_mono (h : G ≤ H) : G.cliqueFinset n ⊆ H.cliqueFinset n := monotone_filter_right _ fun _ ↦ IsNClique.mono h diff --git a/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean b/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean index 3a54a207a63cd..3ee260f7663ee 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean @@ -1077,11 +1077,11 @@ theorem induce_mono (hg : G' ≤ G'') (hs : s ⊆ s') : G'.induce s ≤ G''.indu intro v w hv hw ha exact ⟨hs hv, hs hw, hg.2 ha⟩ -@[mono] +@[gcongr, mono] theorem induce_mono_left (hg : G' ≤ G'') : G'.induce s ≤ G''.induce s := induce_mono hg subset_rfl -@[mono] +@[gcongr, mono] theorem induce_mono_right (hs : s ⊆ s') : G'.induce s ≤ G'.induce s' := induce_mono le_rfl hs @@ -1159,12 +1159,12 @@ theorem deleteVerts_empty : G'.deleteVerts ∅ = G' := by theorem deleteVerts_le : G'.deleteVerts s ≤ G' := by constructor <;> simp [Set.diff_subset] -@[mono] +@[gcongr, mono] theorem deleteVerts_mono {G' G'' : G.Subgraph} (h : G' ≤ G'') : G'.deleteVerts s ≤ G''.deleteVerts s := induce_mono h (Set.diff_subset_diff_left h.1) -@[mono] +@[gcongr, mono] theorem deleteVerts_anti {s s' : Set V} (h : s ⊆ s') : G'.deleteVerts s' ≤ G'.deleteVerts s := induce_mono (le_refl _) (Set.diff_subset_diff_right h) diff --git a/Mathlib/Deprecated/Subgroup.lean b/Mathlib/Deprecated/Subgroup.lean index 745637e6ec963..3842d160c48ac 100644 --- a/Mathlib/Deprecated/Subgroup.lean +++ b/Mathlib/Deprecated/Subgroup.lean @@ -600,6 +600,7 @@ theorem normalClosure_subset_iff {s t : Set G} (ht : IsNormalSubgroup t) : s ⊆ t ↔ normalClosure s ⊆ t := ⟨normalClosure_subset ht, Set.Subset.trans subset_normalClosure⟩ +@[gcongr] theorem normalClosure_mono {s t : Set G} : s ⊆ t → normalClosure s ⊆ normalClosure t := fun h => normalClosure_subset normalClosure.is_normal (Set.Subset.trans h subset_normalClosure) diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean index 76c5638c80a15..b62ed22785418 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean @@ -1250,7 +1250,7 @@ theorem affineSpan_pair_le_of_right_mem {p₁ p₂ p₃ : P} (h : p₁ ∈ line[ variable (k) /-- `affineSpan` is monotone. -/ -@[mono] +@[gcongr, mono] theorem affineSpan_mono {s₁ s₂ : Set P} (h : s₁ ⊆ s₂) : affineSpan k s₁ ≤ affineSpan k s₂ := spanPoints_subset_coe_of_subset_coe (Set.Subset.trans h (subset_affineSpan k _)) diff --git a/Mathlib/LinearAlgebra/Basis/Flag.lean b/Mathlib/LinearAlgebra/Basis/Flag.lean index 4ffe11d49a060..2d138c697c2e9 100644 --- a/Mathlib/LinearAlgebra/Basis/Flag.lean +++ b/Mathlib/LinearAlgebra/Basis/Flag.lean @@ -22,7 +22,8 @@ namespace Basis section Semiring -variable {R M : Type*} [Semiring R] [AddCommMonoid M] [Module R M] {n : ℕ} +variable {R M : Type*} [Semiring R] [AddCommMonoid M] [Module R M] {n : ℕ} {b : Basis (Fin n) R M} + {i j : Fin (n + 1)} /-- The subspace spanned by the first `k` vectors of the basis `b`. -/ def flag (b : Basis (Fin n) R M) (k : Fin (n + 1)) : Submodule R M := @@ -64,6 +65,11 @@ theorem isChain_range_flag (b : Basis (Fin n) R M) : IsChain (· ≤ ·) (range theorem flag_strictMono [Nontrivial R] (b : Basis (Fin n) R M) : StrictMono b.flag := Fin.strictMono_iff_lt_succ.2 fun _ ↦ by simp [flag_succ] +@[gcongr] lemma flag_le_flag (hij : i ≤ j) : b.flag i ≤ b.flag j := flag_mono _ hij + +@[gcongr] +lemma flag_lt_flag [Nontrivial R] (hij : i < j) : b.flag i < b.flag j := flag_strictMono _ hij + end Semiring section CommRing diff --git a/Mathlib/LinearAlgebra/Projectivization/Subspace.lean b/Mathlib/LinearAlgebra/Projectivization/Subspace.lean index 2ae201b95720b..b0a0703432616 100644 --- a/Mathlib/LinearAlgebra/Projectivization/Subspace.lean +++ b/Mathlib/LinearAlgebra/Projectivization/Subspace.lean @@ -154,6 +154,9 @@ span of that set. -/ theorem monotone_span : Monotone (span : Set (ℙ K V) → Subspace K V) := gi.gc.monotone_l +@[gcongr] +lemma span_le_span {s t : Set (ℙ K V)} (hst : s ⊆ t) : span s ≤ span t := monotone_span hst + theorem subset_span_trans {S T U : Set (ℙ K V)} (hST : S ⊆ span T) (hTU : T ⊆ span U) : S ⊆ span U := gi.gc.le_u_l_trans hST hTU diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean index 4c9f13b1bf99b..707d7858aebbe 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean @@ -735,7 +735,7 @@ theorem memℒp_of_bounded [IsFiniteMeasure μ] have hb : ∀ᵐ x ∂μ, f x ≤ b := h.mono fun ω h => h.2 (memℒp_const (max |a| |b|)).mono' hX (by filter_upwards [ha, hb] with x using abs_le_max_abs_abs) -@[mono] +@[gcongr, mono] theorem eLpNorm'_mono_measure (f : α → F) (hμν : ν ≤ μ) (hq : 0 ≤ q) : eLpNorm' f q ν ≤ eLpNorm' f q μ := by simp_rw [eLpNorm'] @@ -745,7 +745,7 @@ theorem eLpNorm'_mono_measure (f : α → F) (hμν : ν ≤ μ) (hq : 0 ≤ q) @[deprecated (since := "2024-07-27")] alias snorm'_mono_measure := eLpNorm'_mono_measure -@[mono] +@[gcongr, mono] theorem eLpNormEssSup_mono_measure (f : α → F) (hμν : ν ≪ μ) : eLpNormEssSup f ν ≤ eLpNormEssSup f μ := by simp_rw [eLpNormEssSup] @@ -754,7 +754,7 @@ theorem eLpNormEssSup_mono_measure (f : α → F) (hμν : ν ≪ μ) : @[deprecated (since := "2024-07-27")] alias snormEssSup_mono_measure := eLpNormEssSup_mono_measure -@[mono] +@[gcongr, mono] theorem eLpNorm_mono_measure (f : α → F) (hμν : ν ≤ μ) : eLpNorm f p ν ≤ eLpNorm f p μ := by by_cases hp0 : p = 0 · simp [hp0] diff --git a/Mathlib/MeasureTheory/Function/SimpleFunc.lean b/Mathlib/MeasureTheory/Function/SimpleFunc.lean index 74809c39984b2..f8ef14327a38d 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFunc.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFunc.lean @@ -665,7 +665,7 @@ theorem mem_image_of_mem_range_restrict {r : β} {s : Set α} {f : α →ₛ β} rw [restrict_of_not_measurable hs] at hr exact (h0 <| eq_zero_of_mem_range_zero hr).elim -@[mono] +@[gcongr, mono] theorem restrict_mono [Preorder β] (s : Set α) {f g : α →ₛ β} (H : f ≤ g) : f.restrict s ≤ g.restrict s := if hs : MeasurableSet s then fun x => by diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index a3a28a6bca1d4..00aeccd87421d 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -1334,7 +1334,7 @@ theorem integral_mono_ae {f g : α → ℝ} (hf : Integrable f μ) (hg : Integra exact setToFun_mono (dominatedFinMeasAdditive_weightedSMul μ) (fun s _ _ => weightedSMul_nonneg s) hf hg h -@[mono] +@[gcongr, mono] theorem integral_mono {f g : α → ℝ} (hf : Integrable f μ) (hg : Integrable g μ) (h : f ≤ g) : ∫ a, f a ∂μ ≤ ∫ a, g a ∂μ := integral_mono_ae hf hg <| Eventually.of_forall h diff --git a/Mathlib/MeasureTheory/Integral/DominatedConvergence.lean b/Mathlib/MeasureTheory/Integral/DominatedConvergence.lean index 9740edb400613..a326a634d81ca 100644 --- a/Mathlib/MeasureTheory/Integral/DominatedConvergence.lean +++ b/Mathlib/MeasureTheory/Integral/DominatedConvergence.lean @@ -589,7 +589,7 @@ theorem continuous_parametric_primitive_of_continuous (uIcc_subset_Icc ⟨a_lt.1.le, lt_b.1.le⟩ ⟨a_lt.2.le, lt_b.2.le⟩) exact Eventually.of_forall this _ ≤ ∫ t in Icc (b₀ - δ) (b₀ + δ), M + 1 ∂μ + ∫ _t in Icc a b, δ ∂μ := by - gcongr + gcongr ?_ + ?_ · apply setIntegral_mono_on · exact (hf.uncurry_left _).norm.integrableOn_Icc · exact continuous_const.integrableOn_Icc diff --git a/Mathlib/MeasureTheory/Integral/Lebesgue.lean b/Mathlib/MeasureTheory/Integral/Lebesgue.lean index aff81c505d96f..e5d534014c897 100644 --- a/Mathlib/MeasureTheory/Integral/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Integral/Lebesgue.lean @@ -78,7 +78,7 @@ theorem SimpleFunc.lintegral_eq_lintegral {m : MeasurableSpace α} (f : α → exact le_antisymm (iSup₂_le fun g hg => lintegral_mono hg <| le_rfl) (le_iSup₂_of_le f le_rfl le_rfl) -@[mono] +@[gcongr, mono] theorem lintegral_mono' {m : MeasurableSpace α} ⦃μ ν : Measure α⦄ (hμν : μ ≤ ν) ⦃f g : α → ℝ≥0∞⦄ (hfg : f ≤ g) : ∫⁻ a, f a ∂μ ≤ ∫⁻ a, g a ∂ν := by rw [lintegral, lintegral] diff --git a/Mathlib/Topology/NhdsSet.lean b/Mathlib/Topology/NhdsSet.lean index 9b361653607e1..f0aa111474378 100644 --- a/Mathlib/Topology/NhdsSet.lean +++ b/Mathlib/Topology/NhdsSet.lean @@ -122,7 +122,7 @@ theorem mem_nhdsSet_empty : s ∈ 𝓝ˢ (∅ : Set X) := by simp @[simp] theorem nhdsSet_univ : 𝓝ˢ (univ : Set X) = ⊤ := by rw [isOpen_univ.nhdsSet_eq, principal_univ] -@[mono] +@[gcongr, mono] theorem nhdsSet_mono (h : s ⊆ t) : 𝓝ˢ s ≤ 𝓝ˢ t := sSup_le_sSup <| image_subset _ h From 1ec3e2a083cc898b366fc001bb52cc88edc774c4 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Thu, 5 Dec 2024 01:28:09 +0000 Subject: [PATCH 520/829] chore: golf `Submodule.mem_of_span_eq_top_of_smul_pow_mem` (#19707) using a recent lemma `Ideal.span_range_pow_eq_top`. --- Mathlib/RingTheory/Ideal/Operations.lean | 12 +++--------- 1 file changed, 3 insertions(+), 9 deletions(-) diff --git a/Mathlib/RingTheory/Ideal/Operations.lean b/Mathlib/RingTheory/Ideal/Operations.lean index cb663d99cc855..a04061874ab5e 100644 --- a/Mathlib/RingTheory/Ideal/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Operations.lean @@ -157,16 +157,10 @@ theorem ideal_span_singleton_smul (r : R) (N : Submodule R M) : submodule `M'` of `x`, we only need to show that `r ^ n • x ∈ M'` for some `n` for each `r : s`. -/ theorem mem_of_span_eq_top_of_smul_pow_mem (M' : Submodule R M) (s : Set R) (hs : Ideal.span s = ⊤) (x : M) (H : ∀ r : s, ∃ n : ℕ, ((r : R) ^ n : R) • x ∈ M') : x ∈ M' := by - obtain ⟨s', hs₁, hs₂⟩ := (Ideal.span_eq_top_iff_finite _).mp hs - replace H : ∀ r : s', ∃ n : ℕ, ((r : R) ^ n : R) • x ∈ M' := fun r => H ⟨_, hs₁ r.2⟩ - choose n₁ n₂ using H - let N := s'.attach.sup n₁ - have hs' := Ideal.span_pow_eq_top (s' : Set R) hs₂ N - apply M'.mem_of_span_top_of_smul_mem _ hs' + choose f hf using H + apply M'.mem_of_span_top_of_smul_mem _ (Ideal.span_range_pow_eq_top s hs f) rintro ⟨_, r, hr, rfl⟩ - convert M'.smul_mem (r ^ (N - n₁ ⟨r, hr⟩)) (n₂ ⟨r, hr⟩) using 1 - simp only [Subtype.coe_mk, smul_smul, ← pow_add] - rw [tsub_add_cancel_of_le (Finset.le_sup (s'.mem_attach _) : n₁ ⟨r, hr⟩ ≤ N)] + exact hf r open Pointwise in @[simp] From 639241321ca03eb192b0db486663b9bf248b7824 Mon Sep 17 00:00:00 2001 From: Tian Chen <27919714+peakpoint@users.noreply.github.com> Date: Thu, 5 Dec 2024 03:56:18 +0000 Subject: [PATCH 521/829] feat(Analysis/Convex/Topology): closure of interior and interior of closure for a convex set (#19170) For a convex set `s` with non-empty interior, `closure (interior s) = closure s` and `interior (closure s) = interior s`. --- Mathlib/Analysis/Convex/Topology.lean | 25 +++++++++++++++++++++++++ 1 file changed, 25 insertions(+) diff --git a/Mathlib/Analysis/Convex/Topology.lean b/Mathlib/Analysis/Convex/Topology.lean index ac81c76f071c7..4fe7a8fcd272e 100644 --- a/Mathlib/Analysis/Convex/Topology.lean +++ b/Mathlib/Analysis/Convex/Topology.lean @@ -268,6 +268,31 @@ protected theorem Convex.strictConvex {s : Set E} (hs : Convex 𝕜 s) end ContinuousConstSMul +section ContinuousSMul + +variable [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] + [TopologicalAddGroup E] [TopologicalSpace 𝕜] [OrderTopology 𝕜] [ContinuousSMul 𝕜 E] + +theorem Convex.closure_interior_eq_closure_of_nonempty_interior {s : Set E} (hs : Convex 𝕜 s) + (hs' : (interior s).Nonempty) : closure (interior s) = closure s := + subset_antisymm (closure_mono interior_subset) + fun _ h ↦ closure_mono (hs.openSegment_interior_closure_subset_interior hs'.choose_spec h) + (segment_subset_closure_openSegment (right_mem_segment ..)) + +theorem Convex.interior_closure_eq_interior_of_nonempty_interior {s : Set E} (hs : Convex 𝕜 s) + (hs' : (interior s).Nonempty) : interior (closure s) = interior s := by + refine subset_antisymm ?_ (interior_mono subset_closure) + intro y hy + rcases hs' with ⟨x, hx⟩ + have h := AffineMap.lineMap_apply_one (k := 𝕜) x y + obtain ⟨t, ht1, ht⟩ := AffineMap.lineMap_continuous.tendsto' _ _ h |>.eventually_mem + (mem_interior_iff_mem_nhds.1 hy) |>.exists_gt + apply hs.openSegment_interior_closure_subset_interior hx ht + nth_rw 1 [← AffineMap.lineMap_apply_zero (k := 𝕜) x y, ← image_openSegment] + exact ⟨1, Ioo_subset_openSegment ⟨zero_lt_one, ht1⟩, h⟩ + +end ContinuousSMul + section TopologicalSpace variable [OrderedSemiring 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] From 1f334b58e1480550e3c4612b9990a9e04a9a4e6d Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 5 Dec 2024 05:05:05 +0000 Subject: [PATCH 522/829] chore: make use of `let` more robust (#19736) We are considering changing the behaviour of term mode let. Mathlib would largely be unscathed; approximately 300 out of the 3000 term mode lets would need to change (either to the existing `letI`, or we might give that a new name). However these two declarations would break, and the changes in this PR, which hopefully are fairly neutral, will avoid the problem. --- Mathlib/Data/Nat/Factors.lean | 2 +- Mathlib/Logic/Equiv/TransferInstance.lean | 7 +++---- 2 files changed, 4 insertions(+), 5 deletions(-) diff --git a/Mathlib/Data/Nat/Factors.lean b/Mathlib/Data/Nat/Factors.lean index 5d4c83169c971..67bf2864fa197 100644 --- a/Mathlib/Data/Nat/Factors.lean +++ b/Mathlib/Data/Nat/Factors.lean @@ -37,7 +37,7 @@ def primeFactorsList : ℕ → List ℕ | k + 2 => let m := minFac (k + 2) m :: primeFactorsList ((k + 2) / m) -decreasing_by show (k + 2) / m < (k + 2); exact factors_lemma +decreasing_by exact factors_lemma @[deprecated (since := "2024-06-14")] alias factors := primeFactorsList diff --git a/Mathlib/Logic/Equiv/TransferInstance.lean b/Mathlib/Logic/Equiv/TransferInstance.lean index e1e5ce7366f08..2b78ff6212187 100644 --- a/Mathlib/Logic/Equiv/TransferInstance.lean +++ b/Mathlib/Logic/Equiv/TransferInstance.lean @@ -621,10 +621,9 @@ protected abbrev algebra (e : α ≃ β) [Semiring β] : simp only [apply_symm_apply, Algebra.mul_smul_comm] lemma algebraMap_def (e : α ≃ β) [Semiring β] [Algebra R β] (r : R) : - let _ := Equiv.semiring e - let _ := Equiv.algebra R e - (algebraMap R α) r = e.symm ((algebraMap R β) r) := by - intros + (@algebraMap R α _ (Equiv.semiring e) (Equiv.algebra R e)) r = e.symm ((algebraMap R β) r) := by + let _ := Equiv.semiring e + let _ := Equiv.algebra R e simp only [Algebra.algebraMap_eq_smul_one] show e.symm (r • e 1) = e.symm (r • 1) simp only [Equiv.one_def, apply_symm_apply] From 6a9d0fe97f9b58ab102aa8507d395816e24d09e9 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Thu, 5 Dec 2024 05:55:22 +0000 Subject: [PATCH 523/829] refactor(SetTheory/Ordinal/Arithmetic): redefine `IsLimit` to `IsSuccLimit` (#19054) This is a temporary stepping stone before outright deprecating `IsLimit`. --- Mathlib/Order/SuccPred/Limit.lean | 9 ++ Mathlib/SetTheory/Cardinal/Arithmetic.lean | 2 +- Mathlib/SetTheory/Cardinal/Cofinality.lean | 16 ++- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 121 ++++++++++-------- Mathlib/SetTheory/Ordinal/Notation.lean | 4 +- Mathlib/SetTheory/Ordinal/Principal.lean | 9 +- Mathlib/SetTheory/Ordinal/Topology.lean | 16 +-- .../Topology/Category/Profinite/Nobeling.lean | 5 +- 8 files changed, 102 insertions(+), 80 deletions(-) diff --git a/Mathlib/Order/SuccPred/Limit.lean b/Mathlib/Order/SuccPred/Limit.lean index 8929a9eebe09a..e23e9f3184fd3 100644 --- a/Mathlib/Order/SuccPred/Limit.lean +++ b/Mathlib/Order/SuccPred/Limit.lean @@ -193,6 +193,9 @@ variable [PartialOrder α] theorem isSuccLimit_iff [OrderBot α] : IsSuccLimit a ↔ a ≠ ⊥ ∧ IsSuccPrelimit a := by rw [IsSuccLimit, isMin_iff_eq_bot] +theorem IsSuccLimit.bot_lt [OrderBot α] (h : IsSuccLimit a) : ⊥ < a := + h.ne_bot.bot_lt + variable [SuccOrder α] theorem isSuccPrelimit_of_succ_ne (h : ∀ b, succ b ≠ a) : IsSuccPrelimit a := fun b hba => @@ -223,6 +226,12 @@ theorem mem_range_succ_or_isSuccPrelimit (a) : a ∈ range (succ : α → α) @[deprecated mem_range_succ_or_isSuccPrelimit (since := "2024-09-05")] alias mem_range_succ_or_isSuccLimit := mem_range_succ_or_isSuccPrelimit +theorem isMin_or_mem_range_succ_or_isSuccLimit (a) : + IsMin a ∨ a ∈ range (succ : α → α) ∨ IsSuccLimit a := by + rw [IsSuccLimit] + have := mem_range_succ_or_isSuccPrelimit a + tauto + theorem isSuccPrelimit_of_succ_lt (H : ∀ a < b, succ a < b) : IsSuccPrelimit b := fun a hab => (H a hab.lt).ne (CovBy.succ_eq hab) diff --git a/Mathlib/SetTheory/Cardinal/Arithmetic.lean b/Mathlib/SetTheory/Cardinal/Arithmetic.lean index 0c071383f0154..8e9e3236e07e2 100644 --- a/Mathlib/SetTheory/Cardinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Cardinal/Arithmetic.lean @@ -82,7 +82,7 @@ theorem mul_eq_self {c : Cardinal} (h : ℵ₀ ≤ c) : c * c = c := by · exact (mul_lt_aleph0 qo qo).trans_le ol · suffices (succ (typein LT.lt (g p))).card < #α from (IH _ this qo).trans_lt this rw [← lt_ord] - apply (isLimit_ord ol).2 + apply (isLimit_ord ol).succ_lt rw [e] apply typein_lt_type diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index df29d51ac8b46..786a594aa7434 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -696,7 +696,7 @@ 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 ⟨S, fun a => - let a' := enum r ⟨_, h.2 _ (typein_lt_type r a)⟩ + let a' := enum r ⟨_, h.succ_lt (typein_lt_type r a)⟩ let ⟨b, h, ab⟩ := H a' ⟨b, h, (IsOrderConnected.conn a b a' <| @@ -834,11 +834,13 @@ theorem isStrongLimit_beth {o : Ordinal} (H : IsSuccPrelimit o) : IsStrongLimit · rw [beth_zero] exact isStrongLimit_aleph0 · refine ⟨beth_ne_zero o, fun a ha => ?_⟩ - rw [beth_limit ⟨h, isSuccPrelimit_iff_succ_lt.1 H⟩] at ha - rcases exists_lt_of_lt_ciSup' ha with ⟨⟨i, hi⟩, ha⟩ - have := power_le_power_left two_ne_zero ha.le - rw [← beth_succ] at this - exact this.trans_lt (beth_lt.2 (H.succ_lt hi)) + rw [beth_limit] at ha + · rcases exists_lt_of_lt_ciSup' ha with ⟨⟨i, hi⟩, ha⟩ + have := power_le_power_left two_ne_zero ha.le + rw [← beth_succ] at this + exact this.trans_lt (beth_lt.2 (H.succ_lt hi)) + · rw [isLimit_iff] + exact ⟨h, H⟩ theorem mk_bounded_subset {α : Type*} (h : ∀ x < #α, (2^x) < #α) {r : α → α → Prop} [IsWellOrder α r] (hr : (#α).ord = type r) : #{ s : Set α // Bounded r s } = #α := by @@ -1140,7 +1142,7 @@ theorem derivFamily_lt_ord_lift {ι : Type u} {f : ι → Ordinal → Ordinal} { rw [derivFamily_succ] exact nfpFamily_lt_ord_lift hω (by rwa [hc.cof_eq]) hf - ((isLimit_ord hc.1).2 _ (hb ((lt_succ b).trans hb'))) + ((isLimit_ord hc.1).succ_lt (hb ((lt_succ b).trans hb'))) | H₃ b hb H => intro hb' -- TODO: generalize the universes of the lemmas in this file so we don't have to rely on bsup diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index cc4daa739e721..b01376da2f420 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -3,10 +3,11 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Floris van Doorn, Violeta Hernández Palacios -/ -import Mathlib.SetTheory.Ordinal.Basic -import Mathlib.Data.Nat.SuccPred import Mathlib.Algebra.GroupWithZero.Divisibility +import Mathlib.Data.Nat.SuccPred +import Mathlib.Order.SuccPred.InitialSeg import Mathlib.SetTheory.Cardinal.UnivLE +import Mathlib.SetTheory.Ordinal.Basic /-! # Ordinal arithmetic @@ -197,33 +198,36 @@ theorem lift_pred (o : Ordinal.{v}) : lift.{u} (pred o) = pred (lift.{u} o) := b TODO: deprecate this in favor of `Order.IsSuccLimit`. -/ def IsLimit (o : Ordinal) : Prop := - o ≠ 0 ∧ ∀ a < o, succ a < o + IsSuccLimit o + +theorem isLimit_iff {o} : IsLimit o ↔ o ≠ 0 ∧ IsSuccPrelimit o := by + simp [IsLimit, IsSuccLimit] theorem IsLimit.isSuccPrelimit {o} (h : IsLimit o) : IsSuccPrelimit o := - isSuccPrelimit_iff_succ_lt.mpr h.2 + IsSuccLimit.isSuccPrelimit h @[deprecated IsLimit.isSuccPrelimit (since := "2024-09-05")] alias IsLimit.isSuccLimit := IsLimit.isSuccPrelimit theorem IsLimit.succ_lt {o a : Ordinal} (h : IsLimit o) : a < o → succ a < o := - h.2 a + IsSuccLimit.succ_lt h theorem isSuccPrelimit_zero : IsSuccPrelimit (0 : Ordinal) := isSuccPrelimit_bot @[deprecated isSuccPrelimit_zero (since := "2024-09-05")] alias isSuccLimit_zero := isSuccPrelimit_zero -theorem not_zero_isLimit : ¬IsLimit 0 - | ⟨h, _⟩ => h rfl +theorem not_zero_isLimit : ¬IsLimit 0 := + not_isSuccLimit_bot -theorem not_succ_isLimit (o) : ¬IsLimit (succ o) - | ⟨_, h⟩ => lt_irrefl _ (h _ (lt_succ o)) +theorem not_succ_isLimit (o) : ¬IsLimit (succ o) := + not_isSuccLimit_succ o theorem not_succ_of_isLimit {o} (h : IsLimit o) : ¬∃ a, o = succ a | ⟨a, e⟩ => not_succ_isLimit a (e ▸ h) theorem succ_lt_of_isLimit {o a : Ordinal} (h : IsLimit o) : succ a < o ↔ a < o := - ⟨(lt_succ a).trans, h.2 _⟩ + IsSuccLimit.succ_lt_iff h theorem le_succ_of_isLimit {o} (h : IsLimit o) {a} : o ≤ succ a ↔ o ≤ a := le_iff_le_iff_lt_iff_lt.2 <| succ_lt_of_isLimit h @@ -238,29 +242,23 @@ theorem lt_limit {o} (h : IsLimit o) {a} : a < o ↔ ∃ x < o, a < x := by @[simp] theorem lift_isLimit (o : Ordinal.{v}) : IsLimit (lift.{u,v} o) ↔ IsLimit o := - and_congr (not_congr <| by simpa only [lift_zero] using @lift_inj o 0) - ⟨fun H a h => (lift_lt.{u,v}).1 <| - by simpa only [lift_succ] using H _ (lift_lt.2 h), fun H a h => by - obtain ⟨a', rfl⟩ := mem_range_lift_of_le h.le - rw [← lift_succ, lift_lt] - exact H a' (lift_lt.1 h)⟩ + liftInitialSeg.isSuccLimit_apply_iff theorem IsLimit.pos {o : Ordinal} (h : IsLimit o) : 0 < o := - lt_of_le_of_ne (Ordinal.zero_le _) h.1.symm + IsSuccLimit.bot_lt h + +theorem IsLimit.ne_zero {o : Ordinal} (h : IsLimit o) : o ≠ 0 := + h.pos.ne' theorem IsLimit.one_lt {o : Ordinal} (h : IsLimit o) : 1 < o := by - simpa only [succ_zero] using h.2 _ h.pos + simpa only [succ_zero] using h.succ_lt h.pos theorem IsLimit.nat_lt {o : Ordinal} (h : IsLimit o) : ∀ n : ℕ, (n : Ordinal) < o | 0 => h.pos - | n + 1 => h.2 _ (IsLimit.nat_lt h n) + | n + 1 => h.succ_lt (IsLimit.nat_lt h n) theorem zero_or_succ_or_limit (o : Ordinal) : o = 0 ∨ (∃ a, o = succ a) ∨ IsLimit o := by - classical - exact if o0 : o = 0 then Or.inl o0 - else - if h : ∃ a, o = succ a then Or.inr (Or.inl h) - else Or.inr <| Or.inr ⟨o0, fun _a => (succ_lt_of_not_succ h).2⟩ + simpa [eq_comm] using isMin_or_mem_range_succ_or_isSuccLimit o theorem isLimit_of_not_succ_of_ne_zero {o : Ordinal} (h : ¬∃ a, o = succ a) (h' : o ≠ 0) : IsLimit o := ((zero_or_succ_or_limit o).resolve_left h').resolve_left h @@ -279,23 +277,24 @@ theorem IsLimit.iSup_Iio {o : Ordinal} (h : IsLimit o) : ⨆ a : Iio o, a.1 = o induction at successor ordinals and at limit ordinals, then it holds for all ordinals. -/ @[elab_as_elim] def limitRecOn {C : Ordinal → Sort*} (o : Ordinal) (H₁ : C 0) (H₂ : ∀ o, C o → C (succ o)) - (H₃ : ∀ o, IsLimit o → (∀ o' < o, C o') → C o) : C o := - SuccOrder.prelimitRecOn o (fun o _ ↦ H₂ o) fun o hl ↦ - if h : o = 0 then fun _ ↦ h ▸ H₁ else H₃ o ⟨h, fun _ ↦ hl.succ_lt⟩ + (H₃ : ∀ o, IsLimit o → (∀ o' < o, C o') → C o) : C o := by + refine SuccOrder.limitRecOn o (fun a ha ↦ ?_) (fun a _ ↦ H₂ a) H₃ + convert H₁ + simpa using ha @[simp] -theorem limitRecOn_zero {C} (H₁ H₂ H₃) : @limitRecOn C 0 H₁ H₂ H₃ = H₁ := by - rw [limitRecOn, SuccOrder.prelimitRecOn_of_isSuccPrelimit _ _ isSuccPrelimit_zero, dif_pos rfl] +theorem limitRecOn_zero {C} (H₁ H₂ H₃) : @limitRecOn C 0 H₁ H₂ H₃ = H₁ := + SuccOrder.limitRecOn_isMin _ _ _ isMin_bot @[simp] theorem limitRecOn_succ {C} (o H₁ H₂ H₃) : - @limitRecOn C (succ o) H₁ H₂ H₃ = H₂ o (@limitRecOn C o H₁ H₂ H₃) := by - rw [limitRecOn, limitRecOn, SuccOrder.prelimitRecOn_succ] + @limitRecOn C (succ o) H₁ H₂ H₃ = H₂ o (@limitRecOn C o H₁ H₂ H₃) := + SuccOrder.limitRecOn_succ .. @[simp] theorem limitRecOn_limit {C} (o H₁ H₂ H₃ h) : - @limitRecOn C o H₁ H₂ H₃ = H₃ o h fun x _h => @limitRecOn C x H₁ H₂ H₃ := by - simp_rw [limitRecOn, SuccOrder.prelimitRecOn_of_isSuccPrelimit _ _ h.isSuccPrelimit, dif_neg h.1] + @limitRecOn C o H₁ H₂ H₃ = H₃ o h fun x _h => @limitRecOn C x H₁ H₂ H₃ := + SuccOrder.limitRecOn_of_isSuccLimit .. /-- Bounded recursion on ordinals. Similar to `limitRecOn`, with the assumption `o < l` added to all cases. The final term's domain is the ordinals below `l`. -/ @@ -347,7 +346,7 @@ alias out_no_max_of_succ_lt := toType_noMax_of_succ_lt theorem bounded_singleton {r : α → α → Prop} [IsWellOrder α r] (hr : (type r).IsLimit) (x) : Bounded r {x} := by - refine ⟨enum r ⟨succ (typein r x), hr.2 _ (typein_lt_type r x)⟩, ?_⟩ + refine ⟨enum r ⟨succ (typein r x), hr.succ_lt (typein_lt_type r x)⟩, ?_⟩ intro b hb rw [mem_singleton_iff.1 hb] nth_rw 1 [← enum_typein r x] @@ -398,7 +397,7 @@ theorem IsNormal.strictMono {f} (H : IsNormal f) : StrictMono f := fun a b => limitRecOn b (Not.elim (not_lt_of_le <| Ordinal.zero_le _)) (fun _b IH h => (lt_or_eq_of_le (le_of_lt_succ h)).elim (fun h => (IH h).trans (H.1 _)) fun e => e ▸ H.1 _) - fun _b l _IH h => lt_of_lt_of_le (H.1 a) ((H.2 _ l _).1 le_rfl _ (l.2 _ h)) + fun _b l _IH h => lt_of_lt_of_le (H.1 a) ((H.2 _ l _).1 le_rfl _ (l.succ_lt h)) theorem IsNormal.monotone {f} (H : IsNormal f) : Monotone f := H.strictMono.monotone @@ -460,10 +459,14 @@ theorem IsNormal.trans {f g} (H₁ : IsNormal f) (H₂ : IsNormal g) : IsNormal ⟨fun _x => H₁.lt_iff.2 (H₂.1 _), fun o l _a => H₁.le_set' (· < o) ⟨0, l.pos⟩ g _ fun _c => H₂.2 _ l _⟩ -theorem IsNormal.isLimit {f} (H : IsNormal f) {o} (l : IsLimit o) : IsLimit (f o) := - ⟨ne_of_gt <| (Ordinal.zero_le _).trans_lt <| H.lt_iff.2 l.pos, fun _ h => - let ⟨_b, h₁, h₂⟩ := (H.limit_lt l).1 h - (succ_le_of_lt h₂).trans_lt (H.lt_iff.2 h₁)⟩ +theorem IsNormal.isLimit {f} (H : IsNormal f) {o} (ho : IsLimit o) : IsLimit (f o) := by + rw [isLimit_iff, isSuccPrelimit_iff_succ_lt] + use (H.lt_iff.2 ho.pos).ne_bot + intro a ha + obtain ⟨b, hb, hab⟩ := (H.limit_lt ho).1 ha + rw [← succ_le_iff] at hab + apply hab.trans_lt + rwa [H.lt_iff] theorem add_le_of_limit {a b c : Ordinal} (h : IsLimit b) : a + b ≤ c ↔ ∀ b' < b, a + b' ≤ c := ⟨fun h _ l => (add_le_add_left l.le _).trans h, fun H => @@ -482,7 +485,7 @@ theorem add_le_of_limit {a b c : Ordinal} (h : IsLimit b) : a + b ≤ c ↔ ∀ · exact irrefl _ (this _) intro x rw [← typein_lt_typein (Sum.Lex r s), typein_enum] - have := H _ (h.2 _ (typein_lt_type s x)) + have := H _ (h.succ_lt (typein_lt_type s x)) rw [add_succ, succ_le_iff] at this refine (RelEmbedding.ofMonotone (fun a => ?_) fun a b => ?_).ordinal_type_le.trans_lt this @@ -566,6 +569,9 @@ protected theorem sub_eq_zero_iff_le {a b : Ordinal} : a - b = 0 ↔ a ≤ b := ⟨fun h => by simpa only [h, add_zero] using le_add_sub a b, fun h => by rwa [← Ordinal.le_zero, sub_le, add_zero]⟩ +protected theorem sub_ne_zero_iff_lt {a b : Ordinal} : a - b ≠ 0 ↔ b < a := by + simpa using Ordinal.sub_eq_zero_iff_le.not + theorem sub_sub (a b c : Ordinal) : a - b - c = a - (b + c) := eq_of_forall_ge_iff fun d => by rw [sub_le, sub_le, sub_le, add_assoc] @@ -587,9 +593,12 @@ theorem lt_add_iff {a b c : Ordinal} (hc : c ≠ 0) : a < b + c ↔ ∃ d < c, a rintro ⟨d, hd, ha⟩ exact ha.trans_lt (add_lt_add_left hd b) -theorem isLimit_sub {a b} (l : IsLimit a) (h : b < a) : IsLimit (a - b) := - ⟨ne_of_gt <| lt_sub.2 <| by rwa [add_zero], fun c h => by - rw [lt_sub, add_succ]; exact l.2 _ (lt_sub.1 h)⟩ +theorem isLimit_sub {a b} (ha : IsLimit a) (h : b < a) : IsLimit (a - b) := by + rw [isLimit_iff, Ordinal.sub_ne_zero_iff_lt, isSuccPrelimit_iff_succ_lt] + refine ⟨h, fun c hc ↦ ?_⟩ + rw [lt_sub] at hc ⊢ + rw [add_succ] + exact ha.succ_lt hc @[deprecated isLimit_sub (since := "2024-10-11")] alias sub_isLimit := isLimit_sub @@ -732,7 +741,7 @@ private theorem mul_le_of_limit_aux {α β r s} [IsWellOrder α r] [IsWellOrder exact irrefl _ (this _ _) intro a b rw [← typein_lt_typein (Prod.Lex s r), typein_enum] - have := H _ (h.2 _ (typein_lt_type s b)) + have := H _ (h.succ_lt (typein_lt_type s b)) rw [mul_succ] at this have := ((add_lt_add_iff_left _).2 (typein_lt_type _ a)).trans_le this refine (RelEmbedding.ofMonotone (fun a => ?_) fun a b => ?_).ordinal_type_le.trans_lt this @@ -1897,7 +1906,7 @@ theorem bsup_eq_blsub_iff_lt_bsup {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max apply lt_blsub, fun h => le_antisymm (bsup_le_blsub f) (blsub_le h)⟩ theorem bsup_eq_blsub_of_lt_succ_limit {o : Ordinal.{u}} (ho : IsLimit o) - {f : ∀ a < o, Ordinal.{max u v}} (hf : ∀ a ha, f a ha < f (succ a) (ho.2 a ha)) : + {f : ∀ a < o, Ordinal.{max u v}} (hf : ∀ a ha, f a ha < f (succ a) (ho.succ_lt ha)) : bsup.{_, v} o f = blsub.{_, v} o f := by rw [bsup_eq_blsub_iff_lt_bsup] exact fun i hi => (hf i hi).trans_le (le_bsup f _ _) @@ -1975,7 +1984,7 @@ theorem blsub_comp {o o' : Ordinal.{max u v}} {f : ∀ a < o, Ordinal.{max u v w theorem IsNormal.bsup_eq {f : Ordinal.{u} → Ordinal.{max u v}} (H : IsNormal f) {o : Ordinal.{u}} (h : IsLimit o) : (Ordinal.bsup.{_, v} o fun x _ => f x) = f o := by - rw [← IsNormal.bsup.{u, u, v} H (fun x _ => x) h.1, bsup_id_limit h.2] + rw [← IsNormal.bsup.{u, u, v} H (fun x _ => x) h.ne_bot, bsup_id_limit fun _ ↦ h.succ_lt] theorem IsNormal.blsub_eq {f : Ordinal.{u} → Ordinal.{max u v}} (H : IsNormal f) {o : Ordinal.{u}} (h : IsLimit o) : (blsub.{_, v} o fun x _ => f x) = f o := by @@ -2360,10 +2369,11 @@ theorem one_lt_omega0 : 1 < ω := by simpa only [Nat.cast_one] using nat_lt_omeg @[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias one_lt_omega := one_lt_omega0 -theorem isLimit_omega0 : IsLimit ω := - ⟨omega0_ne_zero, fun o h => by - let ⟨n, e⟩ := lt_omega0.1 h - rw [e]; exact nat_lt_omega0 (n + 1)⟩ +theorem isLimit_omega0 : IsLimit ω := by + rw [isLimit_iff, isSuccPrelimit_iff_succ_lt] + refine ⟨omega0_ne_zero, fun o h => ?_⟩ + obtain ⟨n, rfl⟩ := lt_omega0.1 h + exact nat_lt_omega0 (n + 1) @[deprecated "No deprecation message was provided." (since := "2024-10-14")] alias omega0_isLimit := isLimit_omega0 @@ -2393,8 +2403,8 @@ theorem sup_natCast : sup Nat.cast = ω := alias sup_nat_cast := sup_natCast 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) + | 0 => h.pos + | n + 1 => h.succ_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 @@ -2403,7 +2413,7 @@ theorem omega0_le_of_isLimit {o} (h : IsLimit o) : ω ≤ o := 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 ⟨fun l => ⟨l.ne_zero, ⟨a / ω, le_antisymm ?_ (mul_div_le _ _)⟩⟩, fun h => ?_⟩ · refine (limit_le l).2 fun x hx => le_of_lt ?_ rw [← div_lt omega0_ne_zero, ← succ_le_iff, le_div omega0_ne_zero, mul_succ, add_le_of_limit isLimit_omega0] @@ -2428,7 +2438,7 @@ theorem add_mul_limit_aux {a b c : Ordinal} (ba : b + a = a) (l : IsLimit c) rw [IH _ h] apply (add_le_add_left _ _).trans · rw [← mul_succ] - exact mul_le_mul_left' (succ_le_of_lt <| l.2 _ h) _ + exact mul_le_mul_left' (succ_le_of_lt <| l.succ_lt h) _ · rw [← ba] exact le_add_right _ _) (mul_le_mul_right' (le_add_right _ _) _) @@ -2495,6 +2505,7 @@ namespace Cardinal open Ordinal theorem isLimit_ord {c} (co : ℵ₀ ≤ c) : (ord c).IsLimit := by + rw [isLimit_iff, isSuccPrelimit_iff_succ_lt] refine ⟨fun h => aleph0_ne_zero ?_, fun a => lt_imp_lt_of_le_imp_le fun h => ?_⟩ · rw [← Ordinal.le_zero, ord_le] at h simpa only [card_zero, nonpos_iff_eq_zero] using co.trans h @@ -2509,7 +2520,7 @@ theorem isLimit_ord {c} (co : ℵ₀ ≤ c) : (ord c).IsLimit := by alias ord_isLimit := isLimit_ord theorem noMaxOrder {c} (h : ℵ₀ ≤ c) : NoMaxOrder c.ord.toType := - toType_noMax_of_succ_lt (isLimit_ord h).2 + toType_noMax_of_succ_lt fun _ ↦ (isLimit_ord h).succ_lt end Cardinal diff --git a/Mathlib/SetTheory/Ordinal/Notation.lean b/Mathlib/SetTheory/Ordinal/Notation.lean index 02806888024eb..52e8d862840e0 100644 --- a/Mathlib/SetTheory/Ordinal/Notation.lean +++ b/Mathlib/SetTheory/Ordinal/Notation.lean @@ -802,8 +802,8 @@ theorem repr_opow_aux₁ {e a} [Ne : NF e] [Na : NF a] {a' : Ordinal} (e0 : repr · apply (mul_le_mul_left' (le_succ b) _).trans rw [← add_one_eq_succ, add_mul_succ _ (one_add_of_omega0_le h), add_one_eq_succ, succ_le_iff, Ordinal.mul_lt_mul_iff_left (Ordinal.pos_iff_ne_zero.2 e0)] - exact isLimit_omega0.2 _ l - · apply (principal_mul_omega0 (isLimit_omega0.2 _ h) l).le.trans + exact isLimit_omega0.succ_lt l + · apply (principal_mul_omega0 (isLimit_omega0.succ_lt h) l).le.trans simpa using mul_le_mul_right' (one_le_iff_ne_zero.2 e0) ω section diff --git a/Mathlib/SetTheory/Ordinal/Principal.lean b/Mathlib/SetTheory/Ordinal/Principal.lean index 194abeabc7c50..5ae52a926bb59 100644 --- a/Mathlib/SetTheory/Ordinal/Principal.lean +++ b/Mathlib/SetTheory/Ordinal/Principal.lean @@ -162,8 +162,9 @@ theorem principal_add_of_le_one (ho : o ≤ 1) : Principal (· + ·) o := by · exact principal_zero · exact principal_add_one -theorem isLimit_of_principal_add (ho₁ : 1 < o) (ho : Principal (· + ·) o) : o.IsLimit := - ⟨ho₁.ne_bot, fun _ ha ↦ ho ha ho₁⟩ +theorem isLimit_of_principal_add (ho₁ : 1 < o) (ho : Principal (· + ·) o) : o.IsLimit := by + rw [isLimit_iff, isSuccPrelimit_iff_succ_lt] + exact ⟨ho₁.ne_bot, fun _ ha ↦ ho ha ho₁⟩ @[deprecated (since := "2024-10-16")] alias principal_add_isLimit := isLimit_of_principal_add @@ -388,7 +389,7 @@ theorem mul_lt_omega0_opow (c0 : 0 < c) (ha : a < ω ^ c) (hb : b < ω) : a * b · rcases ((isNormal_opow 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_omega0] - exact l.2 _ hx + exact l.succ_lt hx @[deprecated (since := "2024-09-30")] alias mul_lt_omega_opow := mul_lt_omega0_opow @@ -457,7 +458,7 @@ theorem mul_eq_opow_log_succ (ha : a ≠ 0) (hb : Principal (· * ·) b) (hb₂ have hbo₀ : b ^ log b a ≠ 0 := Ordinal.pos_iff_ne_zero.1 (opow_pos _ (zero_lt_one.trans hb₁)) apply (mul_le_mul_right' (le_of_lt (lt_mul_succ_div a hbo₀)) c).trans rw [mul_assoc, opow_succ] - refine mul_le_mul_left' (hb (hbl.2 _ ?_) hcb).le _ + refine mul_le_mul_left' (hb (hbl.succ_lt ?_) hcb).le _ rw [div_lt hbo₀, ← opow_succ] exact lt_opow_succ_log_self hb₁ _ · rw [opow_succ] diff --git a/Mathlib/SetTheory/Ordinal/Topology.lean b/Mathlib/SetTheory/Ordinal/Topology.lean index a1dd880e86d62..515769035484d 100644 --- a/Mathlib/SetTheory/Ordinal/Topology.lean +++ b/Mathlib/SetTheory/Ordinal/Topology.lean @@ -37,11 +37,11 @@ instance : TopologicalSpace Ordinal.{u} := Preorder.topology Ordinal.{u} instance : OrderTopology Ordinal.{u} := ⟨rfl⟩ theorem isOpen_singleton_iff : IsOpen ({a} : Set Ordinal) ↔ ¬IsLimit a := by - refine ⟨fun h ⟨h₀, hsucc⟩ => ?_, fun ha => ?_⟩ + refine ⟨fun h ha => ?_, fun ha => ?_⟩ · obtain ⟨b, c, hbc, hbc'⟩ := - (mem_nhds_iff_exists_Ioo_subset' ⟨0, Ordinal.pos_iff_ne_zero.2 h₀⟩ ⟨_, lt_succ a⟩).1 + (mem_nhds_iff_exists_Ioo_subset' ⟨0, ha.pos⟩ ⟨_, lt_succ a⟩).1 (h.mem_nhds rfl) - have hba := hsucc b hbc.1 + have hba := ha.succ_lt hbc.1 exact hba.ne (hbc' ⟨lt_succ b, hba.trans hbc.2⟩) · rcases zero_or_succ_or_limit a with (rfl | ⟨b, rfl⟩ | ha') · rw [← bot_eq_zero, ← Set.Iic_bot, ← Iio_succ] @@ -73,7 +73,7 @@ theorem nhds_eq_pure : 𝓝 a = pure a ↔ ¬IsLimit a := theorem isOpen_iff : IsOpen s ↔ ∀ o ∈ s, IsLimit o → ∃ a < o, Set.Ioo a o ⊆ s := by refine isOpen_iff_mem_nhds.trans <| forall₂_congr fun o ho => ?_ by_cases ho' : IsLimit o - · simp only [(nhdsBasis_Ioc ho'.1).mem_iff, ho', true_implies] + · simp only [(nhdsBasis_Ioc ho'.ne_zero).mem_iff, ho', true_implies] refine exists_congr fun a => and_congr_right fun ha => ?_ simp only [← Set.Ioo_insert_right ha, Set.insert_subset_iff, ho, true_and] · simp [nhds_eq_pure.2 ho', ho, ho'] @@ -216,8 +216,8 @@ theorem isNormal_iff_strictMono_and_continuous (f : Ordinal.{u} → Ordinal.{u}) suffices o ∈ f ⁻¹' Set.Iic a from Set.mem_preimage.1 this rw [mem_iff_iSup_of_isClosed (IsClosed.preimage h' (@isClosed_Iic _ _ _ _ a))] exact - ⟨_, toType_nonempty_iff_ne_zero.2 ho.1, typein (· < ·), fun i => h _ (typein_lt_self i), - sup_typein_limit ho.2⟩ + ⟨_, toType_nonempty_iff_ne_zero.2 ho.ne_zero, typein (· < ·), fun i => h _ (typein_lt_self i), + sup_typein_limit fun _ ↦ ho.succ_lt⟩ theorem enumOrd_isNormal_iff_isClosed (hs : ¬ BddAbove s) : IsNormal (enumOrd s) ↔ IsClosed s := by @@ -237,14 +237,14 @@ theorem enumOrd_isNormal_iff_isClosed (hs : ¬ BddAbove s) : · rw [isClosed_iff_bsup] at h suffices enumOrd s a ≤ bsup.{u, u} a fun b (_ : b < a) => enumOrd s b from this.trans (bsup_le H) - obtain ⟨b, hb⟩ := enumOrd_surjective hs (h ha.1 (fun b _ => enumOrd s b) + obtain ⟨b, hb⟩ := enumOrd_surjective hs (h ha.ne_zero (fun b _ => enumOrd s b) fun b _ => enumOrd_mem hs b) rw [← hb] apply Hs.monotone by_contra! hba apply (Hs (lt_succ b)).not_le rw [hb] - exact le_bsup.{u, u} _ _ (ha.2 _ hba) + exact le_bsup.{u, u} _ _ (ha.succ_lt hba) open Set Filter Set.Notation diff --git a/Mathlib/Topology/Category/Profinite/Nobeling.lean b/Mathlib/Topology/Category/Profinite/Nobeling.lean index f421a36182178..c9673426da74d 100644 --- a/Mathlib/Topology/Category/Profinite/Nobeling.lean +++ b/Mathlib/Topology/Category/Profinite/Nobeling.lean @@ -1055,10 +1055,9 @@ theorem Products.limitOrdinal (l : Products I) : l.isGood (π C (ord I · < o)) ∃ (o' : Ordinal), o' < o ∧ l.isGood (π C (ord I · < o')) := by refine ⟨fun h ↦ ?_, fun ⟨o', ⟨ho', hl⟩⟩ ↦ isGood_mono C (le_of_lt ho') hl⟩ use Finset.sup l.val.toFinset (fun a ↦ Order.succ (ord I a)) - have ha : ⊥ < o := by rw [Ordinal.bot_eq_zero, Ordinal.pos_iff_ne_zero]; exact ho.1 have hslt : Finset.sup l.val.toFinset (fun a ↦ Order.succ (ord I a)) < o := by - simp only [Finset.sup_lt_iff ha, List.mem_toFinset] - exact fun b hb ↦ ho.2 _ (prop_of_isGood C (ord I · < o) h b hb) + simp only [Finset.sup_lt_iff ho.pos, List.mem_toFinset] + exact fun b hb ↦ ho.succ_lt (prop_of_isGood C (ord I · < o) h b hb) refine ⟨hslt, fun he ↦ h ?_⟩ have hlt : ∀ i ∈ l.val, ord I i < Finset.sup l.val.toFinset (fun a ↦ Order.succ (ord I a)) := by intro i hi From 3c028308374366a1efbedf336fdb66ec95309dfb Mon Sep 17 00:00:00 2001 From: grunweg Date: Thu, 5 Dec 2024 06:30:00 +0000 Subject: [PATCH 524/829] perf: micro-optimise some syntax linters (#19733) Perform a few local tweaks - which may or may not improve performance, but surely are more readable. - remove superfluous toArray or toList calls - use `HashSet.ofArray` instead of `ofList` or calling `insert` many times - use `findD` instead of `find?` followed by `getD` - replace `Array.push` in a loop with `Array.append` (currently performance-neutral, but this might change in the future) Inspired by #19730. --- Mathlib/Tactic/Linter/FlexibleLinter.lean | 16 ++-- Mathlib/Tactic/Linter/Multigoal.lean | 6 +- Mathlib/Tactic/Linter/UnusedTactic.lean | 95 ++++++++++++----------- 3 files changed, 57 insertions(+), 60 deletions(-) diff --git a/Mathlib/Tactic/Linter/FlexibleLinter.lean b/Mathlib/Tactic/Linter/FlexibleLinter.lean index fbb790088e980..f4ad7305528d7 100644 --- a/Mathlib/Tactic/Linter/FlexibleLinter.lean +++ b/Mathlib/Tactic/Linter/FlexibleLinter.lean @@ -343,8 +343,7 @@ Otherwise, if an `FVarId` with the same `userName` exists in the new context, us If both of these fail, return `default` (i.e. "fail"). -/ def persistFVars (fv : FVarId) (before after : LocalContext) : FVarId := let ldecl := (before.find? fv).getD default - let name := ldecl.userName - (getFVarIdCandidates fv name after).getD 0 default + (getFVarIdCandidates fv ldecl.userName after).getD 0 default /-- `reallyPersist` converts an array of pairs `(fvar, mvar)` to another array of the same type. -/ def reallyPersist @@ -376,11 +375,11 @@ def flexibleLinter : Linter where run := withSetOptionIn fun _stx => do if (← MonadState.get).messages.hasErrors then return let trees ← getInfoTrees - let x := trees.toList.map (extractCtxAndGoals (fun _ => true)) + let x := trees.map (extractCtxAndGoals (fun _ => true)) -- `stains` records pairs `(location, mvar)`, where -- * `location` is either a hypothesis or the main goal modified by a flexible tactic and -- * `mvar` is the metavariable containing the modified location - let mut stains : Array ((FVarId × MVarId) × (Stained × Syntax)) := .empty + let mut stains : Array ((FVarId × MVarId) × (Stained × Syntax)) := #[] let mut msgs : Array (Syntax × Syntax × Stained) := #[] for d in x do for (s, ctx0, ctx1, mvs0, mvs1) in d do let skind := s.getKind @@ -389,17 +388,14 @@ def flexibleLinter : Linter where run := withSetOptionIn fun _stx => do for d in getStained! s do if shouldStain? then for currMVar1 in mvs1 do - let lctx1 := ((ctx1.decls.find? currMVar1).getD default).lctx + let lctx1 := (ctx1.decls.findD currMVar1 default).lctx let locsAfter := d.toFMVarId currMVar1 lctx1 - - for l in locsAfter do - stains := stains.push (l, (d, s)) - + stains := stains ++ locsAfter.map (fun l ↦ (l, (d, s))) else let stained_in_syntax := if usesGoal? skind then (toStained s).insert d else toStained s if !flexible.contains skind then for currMv0 in mvs0 do - let lctx0 := ((ctx0.decls.find? currMv0).getD default).lctx + let lctx0 := (ctx0.decls.findD currMv0 default).lctx let mut foundFvs : Std.HashSet (FVarId × MVarId):= {} for st in stained_in_syntax do for d in st.toFMVarId currMv0 lctx0 do diff --git a/Mathlib/Tactic/Linter/Multigoal.lean b/Mathlib/Tactic/Linter/Multigoal.lean index e1d838948d7f9..cd119e103387b 100644 --- a/Mathlib/Tactic/Linter/Multigoal.lean +++ b/Mathlib/Tactic/Linter/Multigoal.lean @@ -61,7 +61,7 @@ There is some overlap in scope between `ignoreBranch` and `exclusions`. Tactic combinators like `repeat` or `try` are a mix of both. -/ -abbrev exclusions : Std.HashSet SyntaxNodeKind := .ofList [ +abbrev exclusions : Std.HashSet SyntaxNodeKind := .ofArray #[ -- structuring a proof ``Lean.Parser.Term.cdot, ``cdot, @@ -112,7 +112,7 @@ Reasons for ignoring these tactics include There is some overlap in scope between `exclusions` and `ignoreBranch`. -/ -abbrev ignoreBranch : Std.HashSet SyntaxNodeKind := .ofList [ +abbrev ignoreBranch : Std.HashSet SyntaxNodeKind := .ofArray #[ ``Lean.Parser.Tactic.Conv.conv, `Mathlib.Tactic.Conv.convLHS, `Mathlib.Tactic.Conv.convRHS, @@ -161,7 +161,7 @@ def multiGoalLinter : Linter where run := withSetOptionIn fun _stx ↦ do if (← get).messages.hasErrors then return let trees ← getInfoTrees - for t in trees.toArray do + for t in trees do for (s, before, after, n) in getManyGoals t do let goals (k : Nat) := if k == 1 then f!"1 goal" else f!"{k} goals" let fmt ← Command.liftCoreM diff --git a/Mathlib/Tactic/Linter/UnusedTactic.lean b/Mathlib/Tactic/Linter/UnusedTactic.lean index 81dea423b38b4..01c6ba0915ce9 100644 --- a/Mathlib/Tactic/Linter/UnusedTactic.lean +++ b/Mathlib/Tactic/Linter/UnusedTactic.lean @@ -72,32 +72,33 @@ abbrev M := StateRefT (Std.HashMap String.Range Syntax) IO This can be increased dynamically, using `#allow_unused_tactic`. -/ initialize allowedRef : IO.Ref (Std.HashSet SyntaxNodeKind) ← - IO.mkRef <| Std.HashSet.empty - |>.insert `Mathlib.Tactic.Says.says - |>.insert `Batteries.Tactic.«tacticOn_goal-_=>_» + IO.mkRef <| .ofArray #[ + `Mathlib.Tactic.Says.says, + `Batteries.Tactic.«tacticOn_goal-_=>_», -- attempt to speed up, by ignoring more tactics - |>.insert `by - |>.insert `null - |>.insert `«]» - |>.insert ``Lean.Parser.Term.byTactic - |>.insert ``Lean.Parser.Tactic.tacticSeq - |>.insert ``Lean.Parser.Tactic.tacticSeq1Indented - |>.insert ``Lean.Parser.Tactic.tacticTry_ + `by, + `null, + `«]», + ``Lean.Parser.Term.byTactic, + ``Lean.Parser.Tactic.tacticSeq, + ``Lean.Parser.Tactic.tacticSeq1Indented, + ``Lean.Parser.Tactic.tacticTry_, -- the following `SyntaxNodeKind`s play a role in silencing `test`s - |>.insert ``Lean.Parser.Tactic.guardHyp - |>.insert ``Lean.Parser.Tactic.guardTarget - |>.insert ``Lean.Parser.Tactic.failIfSuccess - |>.insert `Mathlib.Tactic.successIfFailWithMsg - |>.insert `Mathlib.Tactic.failIfNoProgress - |>.insert `Mathlib.Tactic.ExtractGoal.extractGoal - |>.insert `Mathlib.Tactic.Propose.propose' - |>.insert `Lean.Parser.Tactic.traceState - |>.insert `Mathlib.Tactic.tacticMatch_target_ - |>.insert ``Lean.Parser.Tactic.change - |>.insert `change? - |>.insert `«tactic#adaptation_note_» - |>.insert `tacticSleep_heartbeats_ - |>.insert `Mathlib.Tactic.«tacticRename_bvar_→__» + ``Lean.Parser.Tactic.guardHyp, + ``Lean.Parser.Tactic.guardTarget, + ``Lean.Parser.Tactic.failIfSuccess, + `Mathlib.Tactic.successIfFailWithMsg, + `Mathlib.Tactic.failIfNoProgress, + `Mathlib.Tactic.ExtractGoal.extractGoal, + `Mathlib.Tactic.Propose.propose', + `Lean.Parser.Tactic.traceState, + `Mathlib.Tactic.tacticMatch_target_, + ``Lean.Parser.Tactic.change, + `change?, + `«tactic#adaptation_note_», + `tacticSleep_heartbeats_, + `Mathlib.Tactic.«tacticRename_bvar_→__» + ] /-- `#allow_unused_tactic` takes an input a space-separated list of identifiers. These identifiers are then allowed by the unused tactic linter: @@ -119,35 +120,35 @@ A list of blacklisted syntax kinds, which are expected to have subterms that con unevaluated tactics. -/ initialize ignoreTacticKindsRef : IO.Ref NameHashSet ← - IO.mkRef <| Std.HashSet.empty - |>.insert `Mathlib.Tactic.Says.says - |>.insert ``Parser.Term.binderTactic - |>.insert ``Lean.Parser.Term.dynamicQuot - |>.insert ``Lean.Parser.Tactic.quotSeq - |>.insert ``Lean.Parser.Tactic.tacticStop_ - |>.insert ``Lean.Parser.Command.notation - |>.insert ``Lean.Parser.Command.mixfix - |>.insert ``Lean.Parser.Tactic.discharger - |>.insert ``Lean.Parser.Tactic.Conv.conv - |>.insert `Batteries.Tactic.seq_focus - |>.insert `Mathlib.Tactic.Hint.registerHintStx - |>.insert `Mathlib.Tactic.LinearCombination.linearCombination - |>.insert `Mathlib.Tactic.LinearCombination'.linearCombination' - |>.insert `Aesop.Frontend.Parser.addRules - |>.insert `Aesop.Frontend.Parser.aesopTactic - |>.insert `Aesop.Frontend.Parser.aesopTactic? + IO.mkRef <| .ofArray #[ + `Mathlib.Tactic.Says.says, + ``Parser.Term.binderTactic, + ``Lean.Parser.Term.dynamicQuot, + ``Lean.Parser.Tactic.quotSeq, + ``Lean.Parser.Tactic.tacticStop_, + ``Lean.Parser.Command.notation, + ``Lean.Parser.Command.mixfix, + ``Lean.Parser.Tactic.discharger, + ``Lean.Parser.Tactic.Conv.conv, + `Batteries.Tactic.seq_focus, + `Mathlib.Tactic.Hint.registerHintStx, + `Mathlib.Tactic.LinearCombination.linearCombination, + `Mathlib.Tactic.LinearCombination'.linearCombination', + `Aesop.Frontend.Parser.addRules, + `Aesop.Frontend.Parser.aesopTactic, + `Aesop.Frontend.Parser.aesopTactic?, -- the following `SyntaxNodeKind`s play a role in silencing `test`s - |>.insert ``Lean.Parser.Tactic.failIfSuccess - |>.insert `Mathlib.Tactic.successIfFailWithMsg - |>.insert `Mathlib.Tactic.failIfNoProgress + ``Lean.Parser.Tactic.failIfSuccess, + `Mathlib.Tactic.successIfFailWithMsg, + `Mathlib.Tactic.failIfNoProgress + ] /-- Is this a syntax kind that contains intentionally unused tactic subterms? -/ def isIgnoreTacticKind (ignoreTacticKinds : NameHashSet) (k : SyntaxNodeKind) : Bool := k.components.contains `Conv || "slice".isPrefixOf k.toString || - match k with - | .str _ "quot" => true - | _ => ignoreTacticKinds.contains k + k matches .str _ "quot" || + ignoreTacticKinds.contains k /-- Adds a new syntax kind whose children will be ignored by the `unusedTactic` linter. From 06422de9546612e92bafe15d4aa557b4b6a220d1 Mon Sep 17 00:00:00 2001 From: Xavier Roblot Date: Thu, 5 Dec 2024 07:37:53 +0000 Subject: [PATCH 525/829] feat(Analysis/BoxIntegral/UnitPartition): Prove results linking integral point counting and integrals (#12405) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We prove the following result: > Let `s` be a bounded, measurable set of `ι → ℝ` whose frontier has zero volume and let `F` be a > continuous function. Then the limit as `n → ∞` of `∑ F x / n ^ card ι`, where the sum is over the > points in `s ∩ n⁻¹ • (ι → ℤ)`, tends to the integral of `F` over `s`. using Riemann integration. As a special case, we deduce that > The limit as `n → ∞` of `card (s ∩ n⁻¹ • (ι → ℤ)) / n ^ card ι` tends to the volume of `s`. Both of these statements are for a variable `n : ℕ`. However, with the additional hypothesis: `x • s ⊆ y • s` whenever `0 < x ≤ y`, we generalize the previous statement to a real variable. This PR is part of the proof of the Analytic Class Number Formula. Co-authored-by: Xavier Roblot <46200072+xroblot@users.noreply.github.com> --- Mathlib/Algebra/Group/Indicator.lean | 10 +- Mathlib/Algebra/Module/ZLattice/Basic.lean | 15 +- .../Analysis/BoxIntegral/UnitPartition.lean | 197 +++++++++++++++++- Mathlib/LinearAlgebra/Basis/Basic.lean | 4 + Mathlib/Topology/ContinuousOn.lean | 12 ++ 5 files changed, 233 insertions(+), 5 deletions(-) diff --git a/Mathlib/Algebra/Group/Indicator.lean b/Mathlib/Algebra/Group/Indicator.lean index 8dfed5c563fb9..604906c6e8770 100644 --- a/Mathlib/Algebra/Group/Indicator.lean +++ b/Mathlib/Algebra/Group/Indicator.lean @@ -120,9 +120,17 @@ set. -/ theorem mem_of_mulIndicator_ne_one (h : mulIndicator s f a ≠ 1) : a ∈ s := not_imp_comm.1 (fun hn => mulIndicator_of_not_mem hn f) h -@[to_additive] +/-- See `Set.eqOn_mulIndicator'` for the version with `sᶜ`. -/ +@[to_additive + "See `Set.eqOn_indicator'` for the version with `sᶜ`"] theorem eqOn_mulIndicator : EqOn (mulIndicator s f) f s := fun _ hx => mulIndicator_of_mem hx f +/-- See `Set.eqOn_mulIndicator` for the version with `s`. -/ +@[to_additive + "See `Set.eqOn_indicator` for the version with `s`."] +theorem eqOn_mulIndicator' : EqOn (mulIndicator s f) 1 sᶜ := + fun _ hx => mulIndicator_of_not_mem hx f + @[to_additive] theorem mulSupport_mulIndicator_subset : mulSupport (s.mulIndicator f) ⊆ s := fun _ hx => hx.imp_symm fun h => mulIndicator_of_not_mem h f diff --git a/Mathlib/Algebra/Module/ZLattice/Basic.lean b/Mathlib/Algebra/Module/ZLattice/Basic.lean index cadfde60a1179..b8b6b0b57e2f6 100644 --- a/Mathlib/Algebra/Module/ZLattice/Basic.lean +++ b/Mathlib/Algebra/Module/ZLattice/Basic.lean @@ -62,11 +62,17 @@ variable (b : Basis ι K E) theorem span_top : span K (span ℤ (Set.range b) : Set E) = ⊤ := by simp [span_span_of_tower] - theorem map {F : Type*} [NormedAddCommGroup F] [NormedSpace K F] (f : E ≃ₗ[K] F) : Submodule.map (f.restrictScalars ℤ) (span ℤ (Set.range b)) = span ℤ (Set.range (b.map f)) := by simp_rw [Submodule.map_span, LinearEquiv.restrictScalars_apply, Basis.coe_map, Set.range_comp] +open scoped Pointwise in +theorem smul {c : K} (hc : c ≠ 0) : + c • span ℤ (Set.range b) = span ℤ (Set.range (b.isUnitSMul (fun _ ↦ hc.isUnit))) := by + rw [smul_span, Set.smul_set_range] + congr! + rw [Basis.isUnitSMul_apply] + /-- The fundamental domain of the ℤ-lattice spanned by `b`. See `ZSpan.isAddFundamentalDomain` for the proof that it is a fundamental domain. -/ def fundamentalDomain : Set E := {m | ∀ i, b.repr m i ∈ Set.Ico (0 : K) 1} @@ -309,6 +315,13 @@ instance [Finite ι] : DiscreteTopology (span ℤ (Set.range b)) := by instance [Finite ι] : DiscreteTopology (span ℤ (Set.range b)).toAddSubgroup := inferInstanceAs <| DiscreteTopology (span ℤ (Set.range b)) +theorem setFinite_inter [ProperSpace E] [Finite ι] {s : Set E} (hs : Bornology.IsBounded s) : + Set.Finite (s ∩ span ℤ (Set.range b)) := by + have : DiscreteTopology (span ℤ (Set.range b)) := inferInstance + refine Metric.finite_isBounded_inter_isClosed hs ?_ + change IsClosed ((span ℤ (Set.range b)).toAddSubgroup : Set E) + exact AddSubgroup.isClosed_of_discrete + @[measurability] theorem fundamentalDomain_measurableSet [MeasurableSpace E] [OpensMeasurableSpace E] [Finite ι] : MeasurableSet (fundamentalDomain b) := by diff --git a/Mathlib/Analysis/BoxIntegral/UnitPartition.lean b/Mathlib/Analysis/BoxIntegral/UnitPartition.lean index 092ae83e8c61f..9c0e70ce75075 100644 --- a/Mathlib/Analysis/BoxIntegral/UnitPartition.lean +++ b/Mathlib/Analysis/BoxIntegral/UnitPartition.lean @@ -3,6 +3,8 @@ Copyright (c) 2024 Xavier Roblot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Xavier Roblot -/ +import Mathlib.Algebra.Module.ZLattice.Basic +import Mathlib.Analysis.BoxIntegral.Integrability import Mathlib.Analysis.BoxIntegral.Partition.Measure import Mathlib.Analysis.BoxIntegral.Partition.Tagged @@ -11,13 +13,12 @@ import Mathlib.Analysis.BoxIntegral.Partition.Tagged Fix `n` a positive integer. `BoxIntegral.unitPartition.box` are boxes in `ι → ℝ` obtained by dividing the unit box uniformly into boxes of side length `1 / n` and translating the boxes by -the lattice `ι → ℤ` so that they cover the whole space. For fixed `n`, there are indexed by vectors `ν : ι → ℤ`. Let `B` be a `BoxIntegral`. A `unitPartition.box` is admissible for `B` (more precisely its index is admissible) if it is contained in `B`. There are finitely many admissible `unitPartition.box` for `B` and thus we can form the corresponing tagged prepartition, see -`BoxIntegral.unitPartition.prepartition` (note that each `unitPartition.Box` comes with its +`BoxIntegral.unitPartition.prepartition` (note that each `unitPartition.box` comes with its tag situated at its "upper most" vertex). If `B` satisfies `hasIntegralVertices`, that is its vertices are in `ι → ℤ`, then the corresponding prepartition is actually a partition. @@ -35,6 +36,20 @@ coordinates in `ℤ` * `BoxIntegral.unitPartition.prepartition_isPartition`: For `B : BoxIntegral.Box`, if `B` has integral vertices, then the prepartition of `unitPartition.box` admissible for `B` is a partition of `B`. + +* `tendsto_tsum_div_pow_atTop_integral`: let `s` be a bounded, measurable set of `ι → ℝ` +whose frontier has zero volume and let `F` be a continuous function. Then the limit as `n → ∞` +of `∑ F x / n ^ card ι`, where the sum is over the points in `s ∩ n⁻¹ • (ι → ℤ)`, tends to the +integral of `F` over `s`. + +* `tendsto_card_div_pow_atTop_volume`: let `s` be a bounded, measurable set of `ι → ℝ` whose +frontier has zero volume. Then the limit as `n → ∞` of `card (s ∩ n⁻¹ • (ι → ℤ)) / n ^ card ι` +tends to the volume of `s`. + +* `tendsto_card_div_pow_atTop_volume'`: a version of `tendsto_card_div_pow_atTop_volume` where we +assume in addition that `x • s ⊆ y • s` whenever `0 < x ≤ y`. Then we get the same limit +`card (s ∩ x⁻¹ • (ι → ℤ)) / x ^ card ι → volume s` but the limit is over a real variable `x`. + -/ noncomputable section @@ -186,7 +201,7 @@ theorem setFinite_index {s : Set (ι → ℝ)} (hs₁ : NullMeasurableSet s) (hs (by simp only [Set.iUnion_subset_iff, Set.inter_subset_right, implies_true]) hs₂ · rw [Set.mem_setOf, Set.inter_eq_self_of_subset_left hν, volume_box] -/-- For `B : BoxIntegral.Box`, the set of indices of `unitPartition.Box` that are subsets of `B`. +/-- For `B : BoxIntegral.Box`, the set of indices of `unitPartition.box` that are subsets of `B`. This is a finite set. These boxes cover `B` if it has integral vertices, see `unitPartition.prepartition_isPartition`. -/ def admissibleIndex (B : Box ι) : Finset (ι → ℤ) := by @@ -293,4 +308,180 @@ theorem prepartition_isPartition {B : Box ι} (hB : hasIntegralVertices B) : rw [TaggedPrepartition.mem_toPrepartition, mem_prepartition_iff] exact ⟨index n x, mem_admissibleIndex_of_mem_box n hB hx, rfl⟩ +open Submodule Pointwise BigOperators + +open scoped Pointwise + +variable (c : ℝ) (s : Set (ι → ℝ)) (F : (ι → ℝ) → ℝ) + +-- The image of `ι → ℤ` inside `ι → ℝ` +local notation "L" => span ℤ (Set.range (Pi.basisFun ℝ ι)) + +variable {n} in +theorem mem_smul_span_iff {v : ι → ℝ} : + v ∈ (n : ℝ)⁻¹ • L ↔ ∀ i, n * v i ∈ Set.range (algebraMap ℤ ℝ) := by + rw [ZSpan.smul _ (inv_ne_zero (NeZero.ne _)), Basis.mem_span_iff_repr_mem] + simp_rw [Basis.repr_isUnitSMul, Pi.basisFun_repr, Units.smul_def, Units.val_inv_eq_inv_val, + IsUnit.unit_spec, inv_inv, smul_eq_mul] + +theorem tag_mem_smul_span (ν : ι → ℤ) : + tag n ν ∈ (n : ℝ)⁻¹ • L := by + refine mem_smul_span_iff.mpr fun i ↦ ⟨ν i + 1, ?_⟩ + rw [tag_apply, div_eq_inv_mul, ← mul_assoc, mul_inv_cancel_of_invertible, one_mul, map_add, + map_one, eq_intCast] + +theorem tag_index_eq_self_of_mem_smul_span {x : ι → ℝ} (hx : x ∈ (n : ℝ)⁻¹ • L) : + tag n (index n x) = x := by + rw [mem_smul_span_iff] at hx + ext i + obtain ⟨a, ha⟩ : ∃ a : ℤ, a = n * x i := hx i + rwa [tag_apply, index_apply, Int.cast_sub, Int.cast_one, sub_add_cancel, ← ha, Int.ceil_intCast, + div_eq_iff (NeZero.ne _), mul_comm] + +theorem eq_of_mem_smul_span_of_index_eq_index {x y : ι → ℝ} (hx : x ∈ (n : ℝ)⁻¹ • L) + (hy : y ∈ (n : ℝ)⁻¹ • L) (h : index n x = index n y) : x = y := by + rw [← tag_index_eq_self_of_mem_smul_span n hx, ← tag_index_eq_self_of_mem_smul_span n hy, h] + +theorem integralSum_eq_tsum_div {B : Box ι} (hB : hasIntegralVertices B) (hs₀ : s ≤ B) : + integralSum (Set.indicator s F) (BoxAdditiveMap.toSMul (Measure.toBoxAdditive volume)) + (prepartition n B) = (∑' x : ↑(s ∩ (n : ℝ)⁻¹ • L), F x) / n ^ card ι := by + classical + unfold integralSum + have : Fintype ↑(s ∩ (n : ℝ)⁻¹ • L) := by + apply Set.Finite.fintype + rw [← coe_pointwise_smul, ZSpan.smul _ (inv_ne_zero (NeZero.ne _))] + exact ZSpan.setFinite_inter _ (B.isBounded.subset hs₀) + rw [tsum_fintype, Finset.sum_set_coe, Finset.sum_div, eq_comm] + simp_rw [Set.indicator_apply, apply_ite, BoxAdditiveMap.toSMul_apply, Measure.toBoxAdditive_apply, + smul_eq_mul, mul_zero, Finset.sum_ite, Finset.sum_const_zero, add_zero] + refine Finset.sum_bij (fun x _ ↦ box n (index n x)) (fun _ hx ↦ Finset.mem_filter.mpr ?_) + (fun _ hx _ hy h ↦ ?_) (fun I hI ↦ ?_) (fun _ hx ↦ ?_) + · rw [Set.mem_toFinset] at hx + refine ⟨mem_prepartition_boxes_iff.mpr + ⟨index n _, mem_admissibleIndex_of_mem_box n hB (hs₀ hx.1), rfl⟩, ?_⟩ + simp_rw [prepartition_tag n (mem_admissibleIndex_of_mem_box n hB (hs₀ hx.1)), + tag_index_eq_self_of_mem_smul_span n hx.2, hx.1] + · rw [Set.mem_toFinset] at hx hy + exact eq_of_mem_smul_span_of_index_eq_index n hx.2 hy.2 (box_injective n h) + · rw [Finset.mem_filter] at hI + refine ⟨(prepartition n B).tag I, Set.mem_toFinset.mpr ⟨hI.2, ?_⟩, box_index_tag_eq_self n hI.1⟩ + rw [← box_index_tag_eq_self n hI.1, prepartition_tag n + (mem_admissibleIndex_of_mem_box n hB (hs₀ hI.2))] + exact tag_mem_smul_span _ _ + · rw [Set.mem_toFinset] at hx + rw [volume_box, prepartition_tag n (mem_admissibleIndex_of_mem_box n hB (hs₀ hx.1)), + tag_index_eq_self_of_mem_smul_span n hx.2, ENNReal.toReal_div, + ENNReal.one_toReal, ENNReal.toReal_pow, ENNReal.toReal_nat, mul_comm_div, one_mul] + +open Filter + +/-- Let `s` be a bounded, measurable set of `ι → ℝ` whose frontier has zero volume and let `F` +be a continuous function. Then the limit as `n → ∞` of `∑ F x / n ^ card ι`, where the sum is +over the points in `s ∩ n⁻¹ • (ι → ℤ)`, tends to the integral of `F` over `s`. -/ +theorem _root_.tendsto_tsum_div_pow_atTop_integral (hF : Continuous F) (hs₁ : IsBounded s) + (hs₂ : MeasurableSet s) (hs₃ : volume (frontier s) = 0) : + Tendsto (fun n : ℕ ↦ (∑' x : ↑(s ∩ (n : ℝ)⁻¹ • L), F x) / n ^ card ι) + atTop (nhds (∫ x in s, F x)) := by + obtain ⟨B, hB, hs₀⟩ := le_hasIntegralVertices_of_isBounded hs₁ + refine Metric.tendsto_atTop.mpr fun ε hε ↦ ?_ + have h₁ : ∃ C, ∀ x ∈ Box.Icc B, ‖Set.indicator s F x‖ ≤ C := by + obtain ⟨C₀, h₀⟩ := (Box.isCompact_Icc B).exists_bound_of_continuousOn hF.continuousOn + refine ⟨max 0 C₀, fun x hx ↦ ?_⟩ + rw [Set.indicator] + split_ifs with hs + · exact le_max_of_le_right (h₀ x hx) + · exact norm_zero.trans_le <|le_max_left 0 _ + have h₂ : ∀ᵐ x, ContinuousAt (s.indicator F) x := by + filter_upwards [compl_mem_ae_iff.mpr hs₃] with _ h + using (hF.continuousOn).continuousAt_indicator h + obtain ⟨r, hr₁, hr₂⟩ := (hasIntegral_iff.mp <| + AEContinuous.hasBoxIntegral (volume : Measure (ι → ℝ)) h₁ h₂ + IntegrationParams.Riemann) (ε / 2) (half_pos hε) + refine ⟨⌈(r 0 0 : ℝ)⁻¹⌉₊, fun n hn ↦ lt_of_le_of_lt ?_ (half_lt_self_iff.mpr hε)⟩ + have : NeZero n := + ⟨Nat.ne_zero_iff_zero_lt.mpr <| (Nat.ceil_pos.mpr (inv_pos.mpr (r 0 0).prop)).trans_le hn⟩ + rw [← integralSum_eq_tsum_div _ s F hB hs₀, ← Measure.restrict_restrict_of_subset hs₀, + ← integral_indicator hs₂] + refine hr₂ 0 _ ⟨?_, fun _ ↦ ?_, fun h ↦ ?_, fun h ↦ ?_⟩ (prepartition_isPartition _ hB) + · rw [show r 0 = fun _ ↦ r 0 0 from funext_iff.mpr (hr₁ 0 rfl)] + apply prepartition_isSubordinate n B + rw [one_div, inv_le_comm₀ (mod_cast (Nat.pos_of_neZero n)) (r 0 0).prop] + exact le_trans (Nat.le_ceil _) (Nat.cast_le.mpr hn) + · exact prepartition_isHenstock n B + · simp only [IntegrationParams.Riemann, Bool.false_eq_true] at h + · simp only [IntegrationParams.Riemann, Bool.false_eq_true] at h + +/-- Let `s` be a bounded, measurable set of `ι → ℝ` whose frontier has zero volume. Then the limit +as `n → ∞` of `card (s ∩ n⁻¹ • (ι → ℤ)) / n ^ card ι` tends to the volume of `s`. This is a +special case of `tendsto_card_div_pow` with `F = 1`. -/ +theorem _root_.tendsto_card_div_pow_atTop_volume (hs₁ : IsBounded s) + (hs₂ : MeasurableSet s) (hs₃ : volume (frontier s) = 0) : + Tendsto (fun n : ℕ ↦ (Nat.card ↑(s ∩ (n : ℝ)⁻¹ • L) : ℝ) / n ^ card ι) + atTop (nhds (volume s).toReal) := by + convert tendsto_tsum_div_pow_atTop_integral s (fun _ ↦ 1) continuous_const hs₁ hs₂ hs₃ + · rw [tsum_const, nsmul_eq_mul, mul_one, Nat.cast_inj] + · rw [setIntegral_const, smul_eq_mul, mul_one] + +private def tendsto_card_div_pow₁ {c : ℝ} (hc : c ≠ 0) : + ↑(s ∩ c⁻¹ • L) ≃ ↑(c • s ∩ L) := + Equiv.subtypeEquiv (Equiv.smulRight hc) (fun x ↦ by + simp_rw [Set.mem_inter_iff, Equiv.smulRight_apply, Set.smul_mem_smul_set_iff₀ hc, + ← Set.mem_inv_smul_set_iff₀ hc]) + +private theorem tendsto_card_div_pow₂ (hs₁ : IsBounded s) + (hs₄ : ∀ ⦃x y : ℝ⦄, 0 < x → x ≤ y → x • s ⊆ y • s) {x y : ℝ} (hx : 0 < x) (hy : x ≤ y) : + Nat.card ↑(s ∩ x⁻¹ • L) ≤ Nat.card ↑(s ∩ y⁻¹ • L) := by + rw [Nat.card_congr (tendsto_card_div_pow₁ s hx.ne'), + Nat.card_congr (tendsto_card_div_pow₁ s (hx.trans_le hy).ne')] + refine Nat.card_mono ?_ ?_ + · exact ZSpan.setFinite_inter _ (IsBounded.smul₀ hs₁ y) + · exact Set.inter_subset_inter_left _ <| hs₄ hx hy + +private theorem tendsto_card_div_pow₃ (hs₁ : IsBounded s) + (hs₄ : ∀ ⦃x y : ℝ⦄, 0 < x → x ≤ y → x • s ⊆ y • s) : + ∀ᶠ x : ℝ in atTop, (Nat.card ↑(s ∩ (⌊x⌋₊ : ℝ)⁻¹ • L) : ℝ) / x ^ card ι ≤ + (Nat.card ↑(s ∩ x⁻¹ • L) : ℝ) / x ^ card ι := by + filter_upwards [eventually_ge_atTop 1] with x hx + gcongr + exact tendsto_card_div_pow₂ s hs₁ hs₄ (Nat.cast_pos.mpr (Nat.floor_pos.mpr hx)) + (Nat.floor_le (zero_le_one.trans hx)) + +private theorem tendsto_card_div_pow₄ (hs₁ : IsBounded s) + (hs₄ : ∀ ⦃x y : ℝ⦄, 0 < x → x ≤ y → x • s ⊆ y • s) : + ∀ᶠ x : ℝ in atTop, (Nat.card ↑(s ∩ x⁻¹ • L) : ℝ) / x ^ card ι ≤ + (Nat.card ↑(s ∩ (⌈x⌉₊ : ℝ)⁻¹ • L) : ℝ) / x ^ card ι := by + filter_upwards [eventually_gt_atTop 0] with x hx + gcongr + exact tendsto_card_div_pow₂ s hs₁ hs₄ hx (Nat.le_ceil _) + +private theorem tendsto_card_div_pow₅ : + (fun x ↦ (Nat.card ↑(s ∩ (⌊x⌋₊ : ℝ)⁻¹ • L) : ℝ) / ⌊x⌋₊ ^ card ι * (⌊x⌋₊ / x) ^ card ι) + =ᶠ[atTop] (fun x ↦ (Nat.card ↑(s ∩ (⌊x⌋₊ : ℝ)⁻¹ • L) : ℝ) / x ^ card ι) := by + filter_upwards [eventually_ge_atTop 1] with x hx + have : 0 < ⌊x⌋₊ := Nat.floor_pos.mpr hx + rw [div_pow, mul_div, div_mul_cancel₀ _ (by positivity)] + +private theorem tendsto_card_div_pow₆ : + (fun x ↦ (Nat.card ↑(s ∩ (⌈x⌉₊ : ℝ)⁻¹ • L) : ℝ) / ⌈x⌉₊ ^ card ι * (⌈x⌉₊ / x) ^ card ι) + =ᶠ[atTop] (fun x ↦ (Nat.card ↑(s ∩ (⌈x⌉₊ : ℝ)⁻¹ • L) : ℝ) / x ^ card ι) := by + filter_upwards [eventually_ge_atTop 1] with x hx + have : 0 < ⌊x⌋₊ := Nat.floor_pos.mpr hx + rw [div_pow, mul_div, div_mul_cancel₀ _ (by positivity)] + +/-- A version of `tendsto_card_div_pow_atTop_volume` for a real variable. -/ +theorem _root_.tendsto_card_div_pow_atTop_volume' (hs₁ : IsBounded s) + (hs₂ : MeasurableSet s) (hs₃ : volume (frontier s) = 0) + (hs₄ : ∀ ⦃x y : ℝ⦄, 0 < x → x ≤ y → x • s ⊆ y • s) : + Tendsto (fun x : ℝ ↦ (Nat.card ↑(s ∩ x⁻¹ • L) : ℝ) / x ^ card ι) + atTop (nhds (volume s).toReal) := by + rw [show (volume s).toReal = (volume s).toReal * 1 ^ card ι by ring] + refine tendsto_of_tendsto_of_tendsto_of_le_of_le' ?_ ?_ + (tendsto_card_div_pow₃ s hs₁ hs₄) (tendsto_card_div_pow₄ s hs₁ hs₄) + · refine Tendsto.congr' (tendsto_card_div_pow₅ s) (Tendsto.mul ?_ (Tendsto.pow ?_ _)) + · exact Tendsto.comp (tendsto_card_div_pow_atTop_volume s hs₁ hs₂ hs₃) tendsto_nat_floor_atTop + · exact tendsto_nat_floor_div_atTop + · refine Tendsto.congr' (tendsto_card_div_pow₆ s) (Tendsto.mul ?_ (Tendsto.pow ?_ _)) + · exact Tendsto.comp (tendsto_card_div_pow_atTop_volume s hs₁ hs₂ hs₃) tendsto_nat_ceil_atTop + · exact tendsto_nat_ceil_div_atTop + end BoxIntegral.unitPartition diff --git a/Mathlib/LinearAlgebra/Basis/Basic.lean b/Mathlib/LinearAlgebra/Basis/Basic.lean index e3322e7d5c035..662046a8624b8 100644 --- a/Mathlib/LinearAlgebra/Basis/Basic.lean +++ b/Mathlib/LinearAlgebra/Basis/Basic.lean @@ -359,6 +359,10 @@ theorem isUnitSMul_apply {v : Basis ι R M} {w : ι → R} (hw : ∀ i, IsUnit ( v.isUnitSMul hw i = w i • v i := unitsSMul_apply i +theorem repr_isUnitSMul {v : Basis ι R₂ M} {w : ι → R₂} (hw : ∀ i, IsUnit (w i)) (x : M) (i : ι) : + (v.isUnitSMul hw).repr x i = (hw i).unit⁻¹ • v.repr x i := + repr_unitsSMul _ _ _ _ + section Fin /-- Let `b` be a basis for a submodule `N` of `M`. If `y : M` is linear independent of `N` diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index 3e1d97b1dedd0..d65b86c859fa0 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -1416,6 +1416,18 @@ protected lemma Continuous.mulIndicator (hs : ∀ a ∈ frontier s, f a = 1) (hf Continuous (mulIndicator s f) := by classical exact hf.piecewise hs continuous_const +@[to_additive] +theorem ContinuousOn.continuousAt_mulIndicator (hf : ContinuousOn f (interior s)) {x : α} + (hx : x ∉ frontier s) : + ContinuousAt (s.mulIndicator f) x := by + rw [← Set.mem_compl_iff, compl_frontier_eq_union_interior] at hx + obtain h | h := hx + · have hs : interior s ∈ 𝓝 x := mem_interior_iff_mem_nhds.mp (by rwa [interior_interior]) + exact ContinuousAt.congr (hf.continuousAt hs) <| Filter.eventuallyEq_iff_exists_mem.mpr + ⟨interior s, hs, Set.eqOn_mulIndicator.symm.mono interior_subset⟩ + · exact ContinuousAt.congr continuousAt_const <| Filter.eventuallyEq_iff_exists_mem.mpr + ⟨sᶜ, mem_interior_iff_mem_nhds.mp h, Set.eqOn_mulIndicator'.symm⟩ + end Indicator theorem IsOpen.ite' (hs : IsOpen s) (hs' : IsOpen s') From a69e6fafa9665613ef0e905036e4d607b75769e0 Mon Sep 17 00:00:00 2001 From: Mitchell Lee Date: Thu, 5 Dec 2024 08:24:08 +0000 Subject: [PATCH 526/829] feat: continuous image of a filter disjoint from the cocompact filter (#19718) We prove that if a filter is disjoint from the cocompact filter, then so is its image under any continuous function. --- Mathlib/Topology/Compactness/Compact.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/Mathlib/Topology/Compactness/Compact.lean b/Mathlib/Topology/Compactness/Compact.lean index 1149080e64cd9..a28aac3af3af1 100644 --- a/Mathlib/Topology/Compactness/Compact.lean +++ b/Mathlib/Topology/Compactness/Compact.lean @@ -867,6 +867,16 @@ theorem Filter.comap_cocompact_le {f : X → Y} (hf : Continuous f) : refine ⟨f '' t, ht.image hf, ?_⟩ simpa using t.subset_preimage_image f +/-- If a filter is disjoint from the cocompact filter, so is its image under any continuous +function. -/ +theorem disjoint_map_cocompact {g : X → Y} {f : Filter X} (hg : Continuous g) + (hf : Disjoint f (Filter.cocompact X)) : Disjoint (map g f) (Filter.cocompact Y) := by + rw [← Filter.disjoint_comap_iff_map, disjoint_iff_inf_le] + calc + f ⊓ (comap g (cocompact Y)) + _ ≤ f ⊓ Filter.cocompact X := inf_le_inf_left f (Filter.comap_cocompact_le hg) + _ = ⊥ := disjoint_iff.mp hf + theorem isCompact_range [CompactSpace X] {f : X → Y} (hf : Continuous f) : IsCompact (range f) := by rw [← image_univ]; exact isCompact_univ.image hf From 7392ce64217a7167b3598dea6c86ace67c76552e Mon Sep 17 00:00:00 2001 From: Lorenzo Luccioli Date: Thu, 5 Dec 2024 08:47:04 +0000 Subject: [PATCH 527/829] feat: add `mul_min` and `min_mul` (#18926) --- Mathlib/Algebra/Order/Group/Unbundled/Basic.lean | 3 ++- Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean | 10 +++++++++- 2 files changed, 11 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean index 5b90caf6e8964..122d83453934b 100644 --- a/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean @@ -537,7 +537,8 @@ section LE variable [LE α] [MulLeftMono α] {a b c d : α} -set_option linter.docPrime false in +/-- See also `div_le_div_iff` for a version that works for `LinearOrderedSemifield` with +additional assumptions. -/ @[to_additive sub_le_sub_iff] theorem div_le_div_iff' : a / b ≤ c / d ↔ a * d ≤ c * b := by simpa only [div_eq_mul_inv] using mul_inv_le_mul_inv_iff' diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean index 0e9d89e4f69f8..a6577a705c6e3 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean @@ -293,9 +293,17 @@ lemma mul_max [CovariantClass α α (· * ·) (· ≤ ·)] (a b c : α) : a * max b c = max (a * b) (a * c) := mul_left_mono.map_max @[to_additive] -lemma max_mul [CovariantClass α α (swap (· * ·)) (· ≤ ·)] : +lemma max_mul [CovariantClass α α (swap (· * ·)) (· ≤ ·)] (a b c : α) : max a b * c = max (a * c) (b * c) := mul_right_mono.map_max +@[to_additive] +lemma mul_min [CovariantClass α α (· * ·) (· ≤ ·)] (a b c : α) : + a * min b c = min (a * b) (a * c) := mul_left_mono.map_min + +@[to_additive] +lemma min_mul [CovariantClass α α (swap (· * ·)) (· ≤ ·)] (a b c : α) : + min a b * c = min (a * c) (b * c) := mul_right_mono.map_min + @[to_additive] lemma min_lt_max_of_mul_lt_mul [MulLeftMono α] [MulRightMono α] (h : a * b < c * d) : min a b < max c d := by From 2a4850225ef21705c6c28efc71446176c0e9036a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 5 Dec 2024 08:47:05 +0000 Subject: [PATCH 528/829] feat: left division by a nonzero element as a `Perm` (#19458) --- .../Algebra/GroupWithZero/Units/Equiv.lean | 27 +++++++++++++++---- 1 file changed, 22 insertions(+), 5 deletions(-) diff --git a/Mathlib/Algebra/GroupWithZero/Units/Equiv.lean b/Mathlib/Algebra/GroupWithZero/Units/Equiv.lean index c93d58a0a6142..2d7511f6ec174 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Equiv.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Equiv.lean @@ -12,18 +12,20 @@ import Mathlib.Algebra.GroupWithZero.Units.Basic assert_not_exists DenselyOrdered -variable {G₀ : Type*} [GroupWithZero G₀] +variable {G₀ : Type*} + +namespace Equiv +section GroupWithZero +variable [GroupWithZero G₀] /-- In a `GroupWithZero` `G₀`, the unit group `G₀ˣ` is equivalent to the subtype of nonzero elements. -/ -@[simps] def unitsEquivNeZero : G₀ˣ ≃ {a : G₀ // a ≠ 0} where +@[simps] def _root_.unitsEquivNeZero : G₀ˣ ≃ {a : G₀ // a ≠ 0} where toFun a := ⟨a, a.ne_zero⟩ invFun a := Units.mk0 _ a.prop left_inv _ := Units.ext rfl right_inv _ := rfl -namespace Equiv - /-- Left multiplication by a nonzero element in a `GroupWithZero` is a permutation of the underlying type. -/ @[simps! (config := .asFn)] @@ -45,10 +47,25 @@ theorem _root_.mulRight_bijective₀ (a : G₀) (ha : a ≠ 0) : Function.Biject /-- Right division by a nonzero element in a `GroupWithZero` is a permutation of the underlying type. -/ @[simps! (config := { simpRhs := true })] -def divRight₀ (a : G₀) (ha : a ≠ 0) : G₀ ≃ G₀ where +def divRight₀ (a : G₀) (ha : a ≠ 0) : Perm G₀ where toFun := (· / a) invFun := (· * a) left_inv _ := by simp [ha] right_inv _ := by simp [ha] +end GroupWithZero + +section CommGroupWithZero +variable [CommGroupWithZero G₀] + +/-- Left division by a nonzero element in a `CommGroupWithZero` is a permutation of the underlying +type. -/ +@[simps! (config := { simpRhs := true })] +def divLeft₀ (a : G₀) (ha : a ≠ 0) : Perm G₀ where + toFun := (a / ·) + invFun := (a / ·) + left_inv _ := by simp [ha] + right_inv _ := by simp [ha] + +end CommGroupWithZero end Equiv From f296c58dcc57460fef7ac2a1676ef93196509dd0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Thu, 5 Dec 2024 08:47:06 +0000 Subject: [PATCH 529/829] chore(Probability/Kernel): create a Composition folder (#19569) This PR creates a new folder and moves three files: - the `Composition` file becomes `Composition/Basic` - `IntegralCompProd` moves to `Composition/IntegralCompProd` - `MeasureCompProd` moves to `Composition/MeasureCompProd` --- Mathlib.lean | 6 +++--- .../Kernel/{Composition.lean => Composition/Basic.lean} | 0 .../Kernel/{ => Composition}/IntegralCompProd.lean | 2 +- .../Kernel/{ => Composition}/MeasureCompProd.lean | 2 +- Mathlib/Probability/Kernel/Disintegration/Basic.lean | 2 +- Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean | 2 +- Mathlib/Probability/Kernel/Disintegration/Density.lean | 2 +- .../Probability/Kernel/Disintegration/StandardBorel.lean | 2 +- Mathlib/Probability/Kernel/Invariance.lean | 2 +- 9 files changed, 10 insertions(+), 10 deletions(-) rename Mathlib/Probability/Kernel/{Composition.lean => Composition/Basic.lean} (100%) rename Mathlib/Probability/Kernel/{ => Composition}/IntegralCompProd.lean (99%) rename Mathlib/Probability/Kernel/{ => Composition}/MeasureCompProd.lean (99%) diff --git a/Mathlib.lean b/Mathlib.lean index a11769a77bada..f71ab46717573 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4158,7 +4158,9 @@ import Mathlib.Probability.Independence.Kernel import Mathlib.Probability.Independence.ZeroOne import Mathlib.Probability.Integration import Mathlib.Probability.Kernel.Basic -import Mathlib.Probability.Kernel.Composition +import Mathlib.Probability.Kernel.Composition.Basic +import Mathlib.Probability.Kernel.Composition.IntegralCompProd +import Mathlib.Probability.Kernel.Composition.MeasureCompProd import Mathlib.Probability.Kernel.CondDistrib import Mathlib.Probability.Kernel.Condexp import Mathlib.Probability.Kernel.Defs @@ -4171,10 +4173,8 @@ import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes import Mathlib.Probability.Kernel.Disintegration.StandardBorel import Mathlib.Probability.Kernel.Disintegration.Unique import Mathlib.Probability.Kernel.Integral -import Mathlib.Probability.Kernel.IntegralCompProd import Mathlib.Probability.Kernel.Invariance import Mathlib.Probability.Kernel.MeasurableIntegral -import Mathlib.Probability.Kernel.MeasureCompProd import Mathlib.Probability.Kernel.RadonNikodym import Mathlib.Probability.Kernel.WithDensity import Mathlib.Probability.Martingale.Basic diff --git a/Mathlib/Probability/Kernel/Composition.lean b/Mathlib/Probability/Kernel/Composition/Basic.lean similarity index 100% rename from Mathlib/Probability/Kernel/Composition.lean rename to Mathlib/Probability/Kernel/Composition/Basic.lean diff --git a/Mathlib/Probability/Kernel/IntegralCompProd.lean b/Mathlib/Probability/Kernel/Composition/IntegralCompProd.lean similarity index 99% rename from Mathlib/Probability/Kernel/IntegralCompProd.lean rename to Mathlib/Probability/Kernel/Composition/IntegralCompProd.lean index e88e7a0f5f916..0fab071f69bea 100644 --- a/Mathlib/Probability/Kernel/IntegralCompProd.lean +++ b/Mathlib/Probability/Kernel/Composition/IntegralCompProd.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Probability.Kernel.Composition +import Mathlib.Probability.Kernel.Composition.Basic import Mathlib.MeasureTheory.Integral.SetIntegral /-! diff --git a/Mathlib/Probability/Kernel/MeasureCompProd.lean b/Mathlib/Probability/Kernel/Composition/MeasureCompProd.lean similarity index 99% rename from Mathlib/Probability/Kernel/MeasureCompProd.lean rename to Mathlib/Probability/Kernel/Composition/MeasureCompProd.lean index 9a7d41a2b61ca..e15b968b5a1ba 100644 --- a/Mathlib/Probability/Kernel/MeasureCompProd.lean +++ b/Mathlib/Probability/Kernel/Composition/MeasureCompProd.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Probability.Kernel.IntegralCompProd +import Mathlib.Probability.Kernel.Composition.IntegralCompProd /-! # Composition-Product of a measure and a kernel diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 22934beefaf08..8d443a75e791c 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Yaël Dillies, Kin Yau James Wong. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Kin Yau James Wong, Rémy Degenne -/ -import Mathlib.Probability.Kernel.MeasureCompProd +import Mathlib.Probability.Kernel.Composition.MeasureCompProd /-! # Disintegration of measures and kernels diff --git a/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean index 4bdb5515cd2ee..a8558585a52da 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ import Mathlib.MeasureTheory.Function.AEEqOfIntegral -import Mathlib.Probability.Kernel.Composition +import Mathlib.Probability.Kernel.Composition.Basic import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes /-! diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index 94af1027b716b..316c64c2fce93 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Probability.Kernel.Composition +import Mathlib.Probability.Kernel.Composition.Basic import Mathlib.Probability.Martingale.Convergence import Mathlib.Probability.Process.PartitionFiltration diff --git a/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean b/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean index d89d6a2aefef8..9f1b75f64926a 100644 --- a/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Probability.Kernel.MeasureCompProd +import Mathlib.Probability.Kernel.Composition.MeasureCompProd import Mathlib.Probability.Kernel.Disintegration.Basic import Mathlib.Probability.Kernel.Disintegration.CondCDF import Mathlib.Probability.Kernel.Disintegration.Density diff --git a/Mathlib/Probability/Kernel/Invariance.lean b/Mathlib/Probability/Kernel/Invariance.lean index e9d40e85a0dd9..c58725ae3dd90 100644 --- a/Mathlib/Probability/Kernel/Invariance.lean +++ b/Mathlib/Probability/Kernel/Invariance.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Kexing Ying. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kexing Ying -/ -import Mathlib.Probability.Kernel.Composition +import Mathlib.Probability.Kernel.Composition.Basic /-! # Invariance of measures along a kernel From 76a532b4543c50a6e952553cea9f8eff4f4a7bb2 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Thu, 5 Dec 2024 08:47:08 +0000 Subject: [PATCH 530/829] feat(GroupTheory/SpecificGroups/ZGroup): A nilpotent group with cyclic Sylow subgroups is cyclic (#19590) This PR proves that a nilpotent Z-group (a group with cyclic Sylow subgroups) is cyclic. --- .../GroupTheory/SpecificGroups/ZGroup.lean | 25 ++++++++++++++++++- 1 file changed, 24 insertions(+), 1 deletion(-) diff --git a/Mathlib/GroupTheory/SpecificGroups/ZGroup.lean b/Mathlib/GroupTheory/SpecificGroups/ZGroup.lean index 6cb4009aa11c0..f2d850b5a8d40 100644 --- a/Mathlib/GroupTheory/SpecificGroups/ZGroup.lean +++ b/Mathlib/GroupTheory/SpecificGroups/ZGroup.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning -/ import Mathlib.Algebra.Squarefree.Basic -import Mathlib.GroupTheory.Sylow +import Mathlib.GroupTheory.Nilpotent /-! # Z-Groups @@ -30,6 +30,9 @@ variable {G G' f} namespace IsZGroup +instance [IsZGroup G] {p : ℕ} [Fact p.Prime] (P : Sylow p G) : IsCyclic P := + isZGroup p Fact.out P + theorem of_squarefree (hG : Squarefree (Nat.card G)) : IsZGroup G := by have : Finite G := Nat.finite_of_card_ne_zero hG.ne_zero refine ⟨fun p hp P ↦ ?_⟩ @@ -59,4 +62,24 @@ theorem of_surjective [Finite G] [hG : IsZGroup G] (hf : Function.Surjective f) instance [Finite G] [IsZGroup G] (H : Subgroup G) [H.Normal] : IsZGroup (G ⧸ H) := of_surjective (QuotientGroup.mk'_surjective H) +theorem exponent_eq_card [Finite G] [IsZGroup G] : Monoid.exponent G = Nat.card G := by + refine dvd_antisymm Group.exponent_dvd_nat_card ?_ + rw [← Nat.factorization_prime_le_iff_dvd Nat.card_pos.ne' Monoid.exponent_ne_zero_of_finite] + intro p hp + have := Fact.mk hp + let P : Sylow p G := default + rw [← hp.pow_dvd_iff_le_factorization Monoid.exponent_ne_zero_of_finite, + ← P.card_eq_multiplicity, ← (isZGroup p hp P).exponent_eq_card] + exact Monoid.exponent_dvd_of_monoidHom P.1.subtype P.1.subtype_injective + +instance [Finite G] [IsZGroup G] [hG : Group.IsNilpotent G] : IsCyclic G := by + have (p : { x // x ∈ (Nat.card G).primeFactors }) : Fact p.1.Prime := + ⟨Nat.prime_of_mem_primeFactors p.2⟩ + let h (p : { x // x ∈ (Nat.card G).primeFactors }) (P : Sylow p G) : CommGroup P := + IsCyclic.commGroup + obtain ⟨ϕ⟩ := ((isNilpotent_of_finite_tfae (G := G)).out 0 4).mp hG + let _ : CommGroup G := + ⟨fun g h ↦ by rw [← ϕ.symm.injective.eq_iff, map_mul, mul_comm, ← map_mul]⟩ + exact IsCyclic.of_exponent_eq_card exponent_eq_card + end IsZGroup From 967e4d70f3d29231a47f79dfa549b4608eae7f61 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 5 Dec 2024 08:47:09 +0000 Subject: [PATCH 531/829] chore(Finset): rename `Insert.comm` to `insert_comm` (#19610) and add the `DecidablePred Finset.Nontrivial` instance. From GrowthInGroups (LeanCamCombi) --- Mathlib/Data/Finset/Insert.lean | 10 ++++++++-- Mathlib/NumberTheory/FLT/Three.lean | 10 +++++----- 2 files changed, 13 insertions(+), 7 deletions(-) diff --git a/Mathlib/Data/Finset/Insert.lean b/Mathlib/Data/Finset/Insert.lean index 1b93ae1f21cc9..8d5f792fbf340 100644 --- a/Mathlib/Data/Finset/Insert.lean +++ b/Mathlib/Data/Finset/Insert.lean @@ -216,6 +216,10 @@ instance (i : α) : Unique ({i} : Finset α) where @[simp] lemma default_singleton (i : α) : ((default : ({i} : Finset α)) : α) = i := rfl +instance Nontrivial.instDecidablePred [DecidableEq α] : + DecidablePred (Finset.Nontrivial (α := α)) := + inferInstanceAs (DecidablePred fun s ↦ ∃ a ∈ s, ∃ b ∈ s, a ≠ b) + end Singleton /-! ### cons -/ @@ -411,9 +415,11 @@ theorem insert_ne_self : insert a s ≠ s ↔ a ∉ s := theorem pair_eq_singleton (a : α) : ({a, a} : Finset α) = {a} := insert_eq_of_mem <| mem_singleton_self _ -theorem Insert.comm (a b : α) (s : Finset α) : insert a (insert b s) = insert b (insert a s) := +theorem insert_comm (a b : α) (s : Finset α) : insert a (insert b s) = insert b (insert a s) := ext fun x => by simp only [mem_insert, or_left_comm] +@[deprecated (since := "2024-11-29")] alias Insert.comm := insert_comm + @[norm_cast] theorem coe_pair {a b : α} : (({a, b} : Finset α) : Set α) = {a, b} := by ext @@ -424,7 +430,7 @@ theorem coe_eq_pair {s : Finset α} {a b : α} : (s : Set α) = {a, b} ↔ s = { rw [← coe_pair, coe_inj] theorem pair_comm (a b : α) : ({a, b} : Finset α) = {b, a} := - Insert.comm a b ∅ + insert_comm a b ∅ theorem insert_idem (a : α) (s : Finset α) : insert a (insert a s) = insert a s := ext fun x => by simp only [mem_insert, ← or_assoc, or_self_iff] diff --git a/Mathlib/NumberTheory/FLT/Three.lean b/Mathlib/NumberTheory/FLT/Three.lean index 7f70138544592..5ec4e154d21a3 100644 --- a/Mathlib/NumberTheory/FLT/Three.lean +++ b/Mathlib/NumberTheory/FLT/Three.lean @@ -85,7 +85,7 @@ private lemma three_dvd_b_of_dvd_a_of_gcd_eq_one_of_case2 {a b c : ℤ} (ha : a rw [add_comm (a ^ 3), add_assoc, add_comm (a ^ 3), ← add_assoc] at HF refine isCoprime_of_gcd_eq_one_of_FLT ?_ HF convert Hgcd using 2 - rw [Finset.pair_comm, Finset.Insert.comm] + rw [Finset.pair_comm, Finset.insert_comm] by_contra! h3b by_cases h3c : 3 ∣ c · apply h3b @@ -132,15 +132,15 @@ theorem fermatLastTheoremThree_of_three_dvd_only_c rw [← sub_eq_zero, sub_eq_add_neg, ← (show Odd 3 by decide).neg_pow] at hF rcases h1 with (h3a | h3b) | h3c · refine fermatLastTheoremThree_of_dvd_a_of_gcd_eq_one_of_case2 ha h3a ?_ H hF - simp only [← Hgcd, Insert.comm, gcd_insert, gcd_singleton, id_eq, ← abs_eq_normalize, abs_neg] + simp only [← Hgcd, insert_comm, gcd_insert, gcd_singleton, id_eq, ← abs_eq_normalize, abs_neg] · rw [add_comm (a ^ 3)] at hF refine fermatLastTheoremThree_of_dvd_a_of_gcd_eq_one_of_case2 hb h3b ?_ H hF - simp only [← Hgcd, Insert.comm, gcd_insert, gcd_singleton, id_eq, ← abs_eq_normalize, abs_neg] + simp only [← Hgcd, insert_comm, gcd_insert, gcd_singleton, id_eq, ← abs_eq_normalize, abs_neg] · rw [add_comm _ ((-c) ^ 3), ← add_assoc] at hF refine fermatLastTheoremThree_of_dvd_a_of_gcd_eq_one_of_case2 (neg_ne_zero.2 hc) (by simp [h3c]) ?_ H hF - rw [Finset.Insert.comm (-c), Finset.pair_comm (-c) b] - simp only [← Hgcd, Insert.comm, gcd_insert, gcd_singleton, id_eq, ← abs_eq_normalize, abs_neg] + rw [Finset.insert_comm (-c), Finset.pair_comm (-c) b] + simp only [← Hgcd, insert_comm, gcd_insert, gcd_singleton, id_eq, ← abs_eq_normalize, abs_neg] section eisenstein From 7cf71cdadc8d94cb08017524b444cb601730935b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 5 Dec 2024 08:47:10 +0000 Subject: [PATCH 532/829] feat: every chain is contained in a flag (#19615) From LeanCamCombi --- Mathlib/Order/Zorn.lean | 20 ++++++++++++++++++++ 1 file changed, 20 insertions(+) diff --git a/Mathlib/Order/Zorn.lean b/Mathlib/Order/Zorn.lean index b0418fa34a94b..833cc2e53c890 100644 --- a/Mathlib/Order/Zorn.lean +++ b/Mathlib/Order/Zorn.lean @@ -178,3 +178,23 @@ theorem IsChain.exists_maxChain (hc : IsChain r c) : ∃ M, @IsMaxChain _ r M cases' hcs₁ hsy hsz hsseq with h h · exact (hcs₀ hsz).right (h hysy) hzsz hyz · exact (hcs₀ hsy).right hysy (h hzsz) hyz + +/-! ### Flags -/ + +namespace Flag + +variable [Preorder α] {c : Set α} {s : Flag α} {a b : α} + +lemma _root_.IsChain.exists_subset_flag (hc : IsChain (· ≤ ·) c) : ∃ s : Flag α, c ⊆ s := + let ⟨s, hs, hcs⟩ := hc.exists_maxChain; ⟨ofIsMaxChain s hs, hcs⟩ + +lemma exists_mem (a : α) : ∃ s : Flag α, a ∈ s := + let ⟨s, hs⟩ := Set.subsingleton_singleton (a := a).isChain.exists_subset_flag + ⟨s, hs rfl⟩ + +lemma exists_mem_mem (hab : a ≤ b) : ∃ s : Flag α, a ∈ s ∧ b ∈ s := by + simpa [Set.insert_subset_iff] using (IsChain.pair hab).exists_subset_flag + +instance : Nonempty (Flag α) := ⟨.ofIsMaxChain _ maxChain_spec⟩ + +end Flag From ea266590e7c060ee0e5a7fa6a95966304975e019 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Thu, 5 Dec 2024 08:47:12 +0000 Subject: [PATCH 533/829] feat(RingTheory/Algebraic): transitivity of algebraicity (#19660) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Main results: * `Subalgebra.algebraicClosure`: If `R` is a domain and `S` is an arbitrary `R`-algebra, then the elements of `S` that are algebraic over `R` form a subalgebra. (The proofs are due to Matthé van der Lee on [MathOverflow](https://mathoverflow.net/questions/482944/when-do-algebraic-elements-form-a-subalgebra#comment1257824_482944).) * `Algebra.IsAlgebraic.trans'`: If `A/S/R` is a tower of algebras and both `A/S` and `S/R` are algebraic, then `A/R` is also algebraic, provided that `S` has no zero divisors and `algebraMap S A` is injective (so `S` can be regarded as an `R`-subalgebra of `A`). * `Transcendental.extendScalars`: an element of an `R`-algebra that is transcendental over `R` remains transcendental over any algebraic `R`-subalgebra that has no zero divisors. * `IsLocalization.isAlgebraic`: a localization of a nontrivial commutative ring is always algebraic. Used to golf `exists_integral_multiples`. Auxiliary result: * `IsAlgebraic.iff_exists_smul_integral`: If `R` is reduced and `S` is an `R`-algebra with injective `algebraMap`, then an element of `S` is algebraic over `R` iff some `R`-multiple is integral over `R`. --- Mathlib/FieldTheory/AlgebraicClosure.lean | 10 +- Mathlib/RingTheory/Algebraic/Basic.lean | 90 ++++-- Mathlib/RingTheory/Algebraic/Integral.lean | 272 +++++++++++++++++- .../DedekindDomain/IntegralClosure.lean | 24 +- Mathlib/RingTheory/Localization/Integral.lean | 15 +- 5 files changed, 343 insertions(+), 68 deletions(-) diff --git a/Mathlib/FieldTheory/AlgebraicClosure.lean b/Mathlib/FieldTheory/AlgebraicClosure.lean index 9a4a0902bc605..056881d31cde2 100644 --- a/Mathlib/FieldTheory/AlgebraicClosure.lean +++ b/Mathlib/FieldTheory/AlgebraicClosure.lean @@ -38,6 +38,9 @@ def algebraicClosure : IntermediateField F E := variable {F E} +theorem algebraicClosure_toSubalgebra : + (algebraicClosure F E).toSubalgebra = integralClosure F E := rfl + /-- An element is contained in the algebraic closure of `F` in `E` if and only if it is an integral element. -/ theorem mem_algebraicClosure_iff' {x : E} : @@ -100,8 +103,7 @@ variable (F E K) /-- The algebraic closure of `F` in `E` is algebraic over `F`. -/ instance isAlgebraic : Algebra.IsAlgebraic F (algebraicClosure F E) := - ⟨fun x ↦ - isAlgebraic_iff.mpr (IsAlgebraic.isIntegral (mem_algebraicClosure_iff.mp x.2)).isAlgebraic⟩ + ⟨fun x ↦ isAlgebraic_iff.mpr x.2.isAlgebraic⟩ /-- The algebraic closure of `F` in `E` is the integral closure of `F` in `E`. -/ instance isIntegralClosure : IsIntegralClosure (algebraicClosure F E) F E := @@ -109,6 +111,10 @@ instance isIntegralClosure : IsIntegralClosure (algebraicClosure F E) F E := end algebraicClosure +protected theorem Transcendental.algebraicClosure {a : E} (ha : Transcendental F a) : + Transcendental (algebraicClosure F E) a := + ha.extendScalars Subtype.val_injective + variable (F E K) /-- An intermediate field of `E / F` is contained in the algebraic closure of `F` in `E` diff --git a/Mathlib/RingTheory/Algebraic/Basic.lean b/Mathlib/RingTheory/Algebraic/Basic.lean index 2768c2dca2d4d..617f4e02b975c 100644 --- a/Mathlib/RingTheory/Algebraic/Basic.lean +++ b/Mathlib/RingTheory/Algebraic/Basic.lean @@ -5,6 +5,7 @@ Authors: Johan Commelin -/ import Mathlib.Algebra.Polynomial.Expand import Mathlib.Algebra.Polynomial.Roots +import Mathlib.RingTheory.Adjoin.Polynomial import Mathlib.RingTheory.Algebraic.Defs import Mathlib.RingTheory.Polynomial.Tower @@ -31,10 +32,25 @@ theorem is_transcendental_of_subsingleton [Subsingleton R] (x : A) : Transcenden variable {R} -variable (R) in +theorem IsAlgebraic.nontrivial {a : A} (h : IsAlgebraic R a) : Nontrivial R := by + contrapose! h + rw [not_nontrivial_iff_subsingleton] at h + apply is_transcendental_of_subsingleton + +variable (R A) + +theorem Algebra.IsAlgebraic.nontrivial [alg : Algebra.IsAlgebraic R A] : Nontrivial R := + (alg.1 0).nontrivial + +instance (priority := low) Algebra.transcendental_of_subsingleton [Subsingleton R] : + Algebra.Transcendental R A := + ⟨⟨0, is_transcendental_of_subsingleton R 0⟩⟩ + theorem Polynomial.transcendental_X : Transcendental R (X (R := R)) := by simp [transcendental_iff] +variable {R A} + theorem IsAlgebraic.of_aeval {r : A} (f : R[X]) (hf : f.natDegree ≠ 0) (hf' : f.leadingCoeff ∈ nonZeroDivisors R) (H : IsAlgebraic R (aeval r f)) : IsAlgebraic R r := by @@ -93,6 +109,17 @@ theorem transcendental_iff_ker_eq_bot {x : A} : Transcendental R x ↔ RingHom.ker (aeval (R := R) x) = ⊥ := by rw [transcendental_iff_injective, RingHom.injective_iff_ker_eq_bot] +theorem Algebra.isAlgebraic_of_not_injective (h : ¬ Function.Injective (algebraMap R A)) : + Algebra.IsAlgebraic R A where + isAlgebraic a := isAlgebraic_iff_not_injective.mpr + fun inj ↦ h <| by convert inj.comp C_injective; ext; simp + +theorem Algebra.injective_of_transcendental [h : Algebra.Transcendental R A] : + Function.Injective (algebraMap R A) := by + rw [transcendental_iff_not_isAlgebraic] at h + contrapose! h + exact isAlgebraic_of_not_injective h + end section zero_ne_one @@ -129,6 +156,16 @@ theorem isAlgebraic_of_mem_rootSet {R : Type u} {A : Type v} [Field R] [Field A] {p : R[X]} {x : A} (hx : x ∈ p.rootSet A) : IsAlgebraic R x := ⟨p, ne_zero_of_mem_rootSet hx, aeval_eq_zero_of_mem_rootSet hx⟩ +variable (S) in +theorem IsLocalization.isAlgebraic [Nontrivial R] (M : Submonoid R) [IsLocalization M S] : + Algebra.IsAlgebraic R S where + isAlgebraic x := by + obtain rfl | hx := eq_or_ne x 0 + · exact isAlgebraic_zero + have ⟨⟨r, m⟩, h⟩ := surj M x + refine ⟨C m.1 * X - C r, fun eq ↦ hx ?_, by simpa [sub_eq_zero, mul_comm x] using h⟩ + rwa [← eq_mk'_iff_mul_eq, show r = 0 by simpa using congr(coeff $eq 0), mk'_zero] at h + open IsScalarTower protected theorem IsAlgebraic.algebraMap {a : S} : @@ -318,11 +355,9 @@ instance algebra_isAlgebraic_bot_right [Nontrivial R] : end Subalgebra theorem IsAlgebraic.of_pow {r : A} {n : ℕ} (hn : 0 < n) (ht : IsAlgebraic R (r ^ n)) : - IsAlgebraic R r := by - obtain ⟨p, p_nonzero, hp⟩ := ht - refine ⟨Polynomial.expand _ n p, ?_, ?_⟩ - · rwa [Polynomial.expand_ne_zero hn] - · rwa [Polynomial.expand_aeval n p r] + IsAlgebraic R r := + have ⟨p, p_nonzero, hp⟩ := ht + ⟨_, by rwa [expand_ne_zero hn], by rwa [expand_aeval n p r]⟩ theorem Transcendental.pow {r : A} (ht : Transcendental R r) {n : ℕ} (hn : 0 < n) : Transcendental R (r ^ n) := fun ht' ↦ ht <| ht'.of_pow hn @@ -449,8 +484,7 @@ section NoZeroSMulDivisors namespace Algebra.IsAlgebraic -variable [CommRing K] [Field L] -variable [Algebra K L] +variable [CommRing K] [Field L] [Algebra K L] theorem algHom_bijective [NoZeroSMulDivisors K L] [Algebra.IsAlgebraic K L] (f : L →ₐ[K] L) : Function.Bijective f := by @@ -502,42 +536,49 @@ variable {R S : Type*} [CommRing R] [Ring S] [Algebra R S] theorem IsAlgebraic.exists_nonzero_coeff_and_aeval_eq_zero {s : S} (hRs : IsAlgebraic R s) (hs : s ∈ nonZeroDivisors S) : - ∃ (q : Polynomial R), q.coeff 0 ≠ 0 ∧ aeval s q = 0 := by + ∃ q : R[X], q.coeff 0 ≠ 0 ∧ aeval s q = 0 := by obtain ⟨p, hp0, hp⟩ := hRs - obtain ⟨q, hpq, hq⟩ := Polynomial.exists_eq_pow_rootMultiplicity_mul_and_not_dvd p hp0 0 + obtain ⟨q, hpq, hq⟩ := exists_eq_pow_rootMultiplicity_mul_and_not_dvd p hp0 0 simp only [C_0, sub_zero, X_pow_mul, X_dvd_iff] at hpq hq rw [hpq, map_mul, aeval_X_pow] at hp exact ⟨q, hq, (nonZeroDivisors S).pow_mem hs (rootMultiplicity 0 p) (aeval s q) hp⟩ +theorem IsAlgebraic.exists_nonzero_eq_adjoin_mul + {s : S} (hRs : IsAlgebraic R s) (hs : s ∈ nonZeroDivisors S) : + ∃ᵉ (t ∈ Algebra.adjoin R {s}) (r ≠ (0 : R)), s * t = algebraMap R S r := by + have ⟨q, hq0, hq⟩ := hRs.exists_nonzero_coeff_and_aeval_eq_zero hs + have ⟨p, hp⟩ := X_dvd_sub_C (p := q) + refine ⟨aeval s p, aeval_mem_adjoin_singleton _ _, _, neg_ne_zero.mpr hq0, ?_⟩ + apply_fun aeval s at hp + rwa [map_sub, hq, zero_sub, map_mul, aeval_X, aeval_C, ← map_neg, eq_comm] at hp + theorem IsAlgebraic.exists_nonzero_dvd {s : S} (hRs : IsAlgebraic R s) (hs : s ∈ nonZeroDivisors S) : - ∃ r : R, r ≠ 0 ∧ s ∣ (algebraMap R S) r := by + ∃ r : R, r ≠ 0 ∧ s ∣ algebraMap R S r := by obtain ⟨q, hq0, hq⟩ := hRs.exists_nonzero_coeff_and_aeval_eq_zero hs - have key := map_dvd (Polynomial.aeval s) (Polynomial.X_dvd_sub_C (p := q)) - rw [map_sub, hq, zero_sub, dvd_neg, Polynomial.aeval_X, Polynomial.aeval_C] at key + have key := map_dvd (aeval s) (X_dvd_sub_C (p := q)) + rw [map_sub, hq, zero_sub, dvd_neg, aeval_X, aeval_C] at key exact ⟨q.coeff 0, hq0, key⟩ /-- A fraction `(a : S) / (b : S)` can be reduced to `(c : S) / (d : R)`, if `b` is algebraic over `R`. -/ theorem IsAlgebraic.exists_smul_eq_mul (a : S) {b : S} (hRb : IsAlgebraic R b) (hb : b ∈ nonZeroDivisors S) : - ∃ᵉ (c : S) (d ≠ (0 : R)), d • a = b * c := by - obtain ⟨r, hr, s, h⟩ := IsAlgebraic.exists_nonzero_dvd (R := R) (S := S) hRb hb - exact ⟨s * a, r, hr, by rw [Algebra.smul_def, h, mul_assoc]⟩ + ∃ᵉ (c : S) (d ≠ (0 : R)), d • a = b * c := + have ⟨r, hr, s, h⟩ := hRb.exists_nonzero_dvd hb + ⟨s * a, r, hr, by rw [Algebra.smul_def, h, mul_assoc]⟩ variable (R) /-- A fraction `(a : S) / (b : S)` can be reduced to `(c : S) / (d : R)`, if `b` is algebraic over `R`. -/ -theorem Algebra.IsAlgebraic.exists_smul_eq_mul [IsDomain S] [Algebra.IsAlgebraic R S] +theorem Algebra.IsAlgebraic.exists_smul_eq_mul [NoZeroDivisors S] [Algebra.IsAlgebraic R S] (a : S) {b : S} (hb : b ≠ 0) : ∃ᵉ (c : S) (d ≠ (0 : R)), d • a = b * c := - (Algebra.IsAlgebraic.isAlgebraic b).exists_smul_eq_mul a (mem_nonZeroDivisors_iff_ne_zero.mpr hb) + (isAlgebraic b).exists_smul_eq_mul a (mem_nonZeroDivisors_of_ne_zero hb) end -variable {R S : Type*} [CommRing R] [CommRing S] - section Field variable {K L : Type*} [Field K] [Field L] [Algebra K L] (A : Subalgebra K L) @@ -601,12 +642,13 @@ end Field section Infinite -theorem Transcendental.infinite {R A : Type*} [CommRing R] [Ring A] [Algebra R A] - [Nontrivial R] {x : A} (hx : Transcendental R x) : Infinite A := +variable {R A : Type*} [CommRing R] [Ring A] [Algebra R A] [Nontrivial R] + +theorem Transcendental.infinite {x : A} (hx : Transcendental R x) : Infinite A := .of_injective _ (transcendental_iff_injective.mp hx) -theorem Algebra.Transcendental.infinite (R A : Type*) [CommRing R] [Ring A] [Algebra R A] - [Nontrivial R] [Algebra.Transcendental R A] : Infinite A := +variable (R A) in +theorem Algebra.Transcendental.infinite [Algebra.Transcendental R A] : Infinite A := have ⟨x, hx⟩ := ‹Algebra.Transcendental R A› hx.infinite diff --git a/Mathlib/RingTheory/Algebraic/Integral.lean b/Mathlib/RingTheory/Algebraic/Integral.lean index b45ecf7023b2f..5a9c3fa6e04f6 100644 --- a/Mathlib/RingTheory/Algebraic/Integral.lean +++ b/Mathlib/RingTheory/Algebraic/Integral.lean @@ -19,7 +19,17 @@ is algebraic and that every algebraic element over a field is integral. over a field. * `IsAlgebraic.of_finite`, `Algebra.IsAlgebraic.of_finite`: finite-dimensional (as module) implies algebraic. -* `exists_integral_multiple`: an algebraic element has a multiple which is integral +* `IsAlgebraic.exists_integral_multiple`: an algebraic element has a multiple which is integral +* `IsAlgebraic.iff_exists_smul_integral`: If `R` is reduced and `S` is an `R`-algebra with + injective `algebraMap`, then an element of `S` is algebraic over `R` iff some `R`-multiple + is integral over `R`. +* `Algebra.IsAlgebraic.trans'`: If `A/S/R` is a tower of algebras and both `A/S` and `S/R` are + algebraic, then `A/R` is also algebraic, provided that `S` has no zero divisors and + `algebraMap S A` is injective (so `S` can be regarded as an `R`-subalgebra of `A`). +* `Subalgebra.algebraicClosure`: If `R` is a domain and `S` is an arbitrary `R`-algebra, + then the elements of `S` that are algebraic over `R` form a subalgebra. +* `Transcendental.extendScalars`: an element of an `R`-algebra that is transcendental over `R` + remains transcendental over any algebraic `R`-subalgebra that has no zero divisors. -/ assert_not_exists LocalRing @@ -66,11 +76,11 @@ alias ⟨IsAlgebraic.isIntegral, _⟩ := isAlgebraic_iff_isIntegral protected instance Algebra.IsAlgebraic.isIntegral [Algebra.IsAlgebraic K A] : Algebra.IsIntegral K A := Algebra.isAlgebraic_iff_isIntegral.mp ‹_› -variable (K) in -theorem Algebra.IsAlgebraic.of_isIntegralClosure (B C : Type*) - [CommRing B] [CommRing C] [Algebra K B] [Algebra K C] [Algebra B C] - [IsScalarTower K B C] [IsIntegralClosure B K C] : Algebra.IsAlgebraic K B := - Algebra.isAlgebraic_iff_isIntegral.mpr (IsIntegralClosure.isIntegral_algebra K C) +theorem Algebra.IsAlgebraic.of_isIntegralClosure (R B C : Type*) [CommRing R] [Nontrivial R] + [CommRing B] [CommRing C] [Algebra R B] [Algebra R C] [Algebra B C] + [IsScalarTower R B C] [IsIntegralClosure B R C] : Algebra.IsAlgebraic R B := + have := IsIntegralClosure.isIntegral_algebra R (A := B) C + inferInstance end Field @@ -146,15 +156,249 @@ end Field end -variable {R S : Type*} [CommRing R] [CommRing S] +variable {R S A : Type*} [CommRing R] [CommRing S] [Ring A] +variable [Algebra R S] [Algebra R A] [Algebra S A] [IsScalarTower R S A] +variable {z : A} {z' : S} -theorem exists_integral_multiple [Algebra R S] {z : S} (hz : IsAlgebraic R z) - (inj : ∀ x, algebraMap R S x = 0 → x = 0) : - ∃ᵉ (x : integralClosure R S) (y ≠ (0 : R)), algebraMap R S y * z = x := by - rcases hz with ⟨p, p_ne_zero, px⟩ +namespace IsAlgebraic + +theorem exists_integral_multiple (hz : IsAlgebraic R z) + (inj : Function.Injective (algebraMap R A)) : + ∃ y ≠ (0 : R), IsIntegral R (y • z) := by + have ⟨p, p_ne_zero, px⟩ := hz set a := p.leadingCoeff have a_ne_zero : a ≠ 0 := mt Polynomial.leadingCoeff_eq_zero.mp p_ne_zero - have x_integral : IsIntegral R (algebraMap R S a * z) := + have x_integral : IsIntegral R (algebraMap R A a * z) := ⟨p.integralNormalization, monic_integralNormalization p_ne_zero, - integralNormalization_aeval_eq_zero px inj⟩ - exact ⟨⟨_, x_integral⟩, a, a_ne_zero, rfl⟩ + integralNormalization_aeval_eq_zero px fun _ ↦ (map_eq_zero_iff _ inj).mp⟩ + exact ⟨_, a_ne_zero, Algebra.smul_def a z ▸ x_integral⟩ + +@[deprecated (since := "2024-11-30")] +alias _root_.exists_integral_multiple := exists_integral_multiple + +theorem _root_.Algebra.IsAlgebraic.exists_integral_multiples [NoZeroDivisors R] + [alg : Algebra.IsAlgebraic R A] (inj : Function.Injective (algebraMap R A)) (s : Finset A) : + ∃ y ≠ (0 : R), ∀ z ∈ s, IsIntegral R (y • z) := by + have := Algebra.IsAlgebraic.nontrivial R A + choose r hr int using fun x ↦ (alg.1 x).exists_integral_multiple inj + refine ⟨∏ x ∈ s, r x, Finset.prod_ne_zero_iff.mpr fun _ _ ↦ hr _, fun _ h ↦ ?_⟩ + classical rw [← Finset.prod_erase_mul _ _ h, mul_smul] + exact (int _).smul _ + +theorem of_smul_isIntegral {y : R} (hy : ¬ IsNilpotent y) + (h : IsIntegral R (y • z)) : IsAlgebraic R z := by + have ⟨p, monic, eval0⟩ := h + refine ⟨p.comp (C y * X), fun h ↦ ?_, by simpa [aeval_comp, Algebra.smul_def] using eval0⟩ + apply_fun (coeff · p.natDegree) at h + have hy0 : y ≠ 0 := by rintro rfl; exact hy .zero + rw [coeff_zero, ← mul_one p.natDegree, ← natDegree_C_mul_X y hy0, + coeff_comp_degree_mul_degree, monic, one_mul, leadingCoeff_C_mul_X] at h + · exact hy ⟨_, h⟩ + · rw [natDegree_C_mul_X _ hy0]; rintro ⟨⟩ + +theorem of_smul {y : R} (hy : y ∈ nonZeroDivisors R) + (h : IsAlgebraic R (y • z)) : IsAlgebraic R z := + have ⟨p, hp, eval0⟩ := h + ⟨_, mt (comp_C_mul_X_eq_zero_iff hy).mp hp, by simpa [aeval_comp, Algebra.smul_def] using eval0⟩ + +theorem iff_exists_smul_integral [IsReduced R] (inj : Function.Injective (algebraMap R A)) : + IsAlgebraic R z ↔ ∃ y ≠ (0 : R), IsIntegral R (y • z) := + ⟨(exists_integral_multiple · inj), fun ⟨_, hy, int⟩ ↦ + of_smul_isIntegral (by rwa [isNilpotent_iff_eq_zero]) int⟩ + +section trans + +variable (R) [NoZeroDivisors S] (inj : Function.Injective (algebraMap S A)) +include inj + +theorem restrictScalars_of_isIntegral [int : Algebra.IsIntegral R S] + {a : A} (h : IsAlgebraic S a) : IsAlgebraic R a := by + by_cases hRS : Function.Injective (algebraMap R S) + on_goal 2 => exact (Algebra.isAlgebraic_of_not_injective + fun h ↦ hRS <| .of_comp (IsScalarTower.algebraMap_eq R S A ▸ h)).1 _ + have := hRS.noZeroDivisors _ (map_zero _) (map_mul _) + have ⟨s, hs, int_s⟩ := h.exists_integral_multiple inj + cases subsingleton_or_nontrivial R + · have := Module.subsingleton R S + exact (is_transcendental_of_subsingleton _ _ h).elim + have ⟨r, hr, _, e⟩ := (int.1 s).isAlgebraic.exists_nonzero_dvd (mem_nonZeroDivisors_of_ne_zero hs) + refine .of_smul_isIntegral (y := r) (by rwa [isNilpotent_iff_eq_zero]) ?_ + rw [Algebra.smul_def, IsScalarTower.algebraMap_apply R S, + e, ← Algebra.smul_def, mul_comm, mul_smul] + exact isIntegral_trans _ (int_s.smul _) + +theorem restrictScalars [Algebra.IsAlgebraic R S] + {a : A} (h : IsAlgebraic S a) : IsAlgebraic R a := by + have ⟨p, hp, eval0⟩ := h + by_cases hRS : Function.Injective (algebraMap R S) + on_goal 2 => exact (Algebra.isAlgebraic_of_not_injective + fun h ↦ hRS <| .of_comp (IsScalarTower.algebraMap_eq R S A ▸ h)).1 _ + have := hRS.noZeroDivisors _ (map_zero _) (map_mul _) + classical + have ⟨r, hr, int⟩ := Algebra.IsAlgebraic.exists_integral_multiples hRS (p.support.image (coeff p)) + let p := (r • p).toSubring (integralClosure R S).toSubring fun s hs ↦ by + obtain ⟨n, hn, rfl⟩ := mem_coeffs_iff.mp hs + exact int _ (Finset.mem_image_of_mem _ <| support_smul _ _ hn) + have : IsAlgebraic (integralClosure R S) a := by + refine ⟨p, ?_, ?_⟩ + · have := NoZeroSMulDivisors.of_algebraMap_injective hRS + simpa only [← Polynomial.map_ne_zero_iff (f := Subring.subtype _) Subtype.val_injective, + p, map_toSubring, smul_ne_zero_iff] using And.intro hr hp + rw [← eval_map_algebraMap, Subalgebra.algebraMap_eq, ← map_map, ← Subalgebra.toSubring_subtype, + map_toSubring, eval_map_algebraMap, ← AlgHom.restrictScalars_apply R, + map_smul, AlgHom.restrictScalars_apply, eval0, smul_zero] + exact restrictScalars_of_isIntegral _ (by exact inj.comp Subtype.val_injective) this + +theorem _root_.IsIntegral.trans_isAlgebraic [alg : Algebra.IsAlgebraic R S] + {a : A} (h : IsIntegral S a) : IsAlgebraic R a := by + cases subsingleton_or_nontrivial A + · have := Algebra.IsAlgebraic.nontrivial R S + exact Subsingleton.elim a 0 ▸ isAlgebraic_zero + · have := Module.nontrivial S A + exact h.isAlgebraic.restrictScalars _ inj + +end trans + +variable [nzd : NoZeroDivisors R] {a b : S} (ha : IsAlgebraic R a) (hb : IsAlgebraic R b) +include ha +omit nzd + +protected lemma neg : IsAlgebraic R (-a) := + have ⟨p, h, eval0⟩ := ha + ⟨algEquivAevalNegX p, EmbeddingLike.map_ne_zero_iff.mpr h, by simpa [← comp_eq_aeval, aeval_comp]⟩ + +protected lemma smul (r : R) : IsAlgebraic R (r • a) := + have ⟨_, hp, eval0⟩ := ha + ⟨_, scaleRoots_ne_zero hp r, Algebra.smul_def r a ▸ scaleRoots_aeval_eq_zero eval0⟩ + +protected lemma nsmul (n : ℕ) : IsAlgebraic R (n • a) := + Nat.cast_smul_eq_nsmul R n a ▸ ha.smul _ + +protected lemma zsmul (n : ℤ) : IsAlgebraic R (n • a) := + Int.cast_smul_eq_zsmul R n a ▸ ha.smul _ + +include hb nzd + +protected lemma mul : IsAlgebraic R (a * b) := by + refine (em _).elim (fun h ↦ ?_) fun h ↦ (Algebra.isAlgebraic_of_not_injective h).1 _ + have ⟨ra, a0, int_a⟩ := ha.exists_integral_multiple h + have ⟨rb, b0, int_b⟩ := hb.exists_integral_multiple h + refine (IsAlgebraic.iff_exists_smul_integral h).mpr ⟨_, mul_ne_zero a0 b0, ?_⟩ + simp_rw [Algebra.smul_def, map_mul, mul_mul_mul_comm, ← Algebra.smul_def] + exact int_a.mul int_b + +protected lemma add : IsAlgebraic R (a + b) := by + refine (em _).elim (fun h ↦ ?_) fun h ↦ (Algebra.isAlgebraic_of_not_injective h).1 _ + have ⟨ra, a0, int_a⟩ := ha.exists_integral_multiple h + have ⟨rb, b0, int_b⟩ := hb.exists_integral_multiple h + refine (IsAlgebraic.iff_exists_smul_integral h).mpr ⟨_, mul_ne_zero b0 a0, ?_⟩ + rw [smul_add, mul_smul, mul_comm, mul_smul] + exact (int_a.smul _).add (int_b.smul _) + +protected lemma sub : IsAlgebraic R (a - b) := + sub_eq_add_neg a b ▸ ha.add hb.neg + +omit hb +protected lemma pow (n : ℕ) : IsAlgebraic R (a ^ n) := + have := ha.nontrivial + n.rec (pow_zero a ▸ isAlgebraic_one) fun _ h ↦ pow_succ a _ ▸ h.mul ha + +end IsAlgebraic + +namespace Algebra + +variable (R) [NoZeroDivisors S] (inj : Function.Injective (algebraMap S A)) +include inj + +/-- Transitivity of algebraicity for algebras over domains. -/ +theorem IsAlgebraic.trans' [Algebra.IsAlgebraic R S] [alg : Algebra.IsAlgebraic S A] : + Algebra.IsAlgebraic R A := + ⟨fun _ ↦ (alg.1 _).restrictScalars _ inj⟩ + +theorem IsIntegral.trans_isAlgebraic [Algebra.IsIntegral R S] [alg : Algebra.IsAlgebraic S A] : + Algebra.IsAlgebraic R A := + ⟨fun _ ↦ (alg.1 _).restrictScalars_of_isIntegral _ inj⟩ + +theorem IsAlgebraic.trans_isIntegral [Algebra.IsAlgebraic R S] [int : Algebra.IsIntegral S A] : + Algebra.IsAlgebraic R A := + ⟨fun _ ↦ (int.1 _).trans_isAlgebraic _ inj⟩ + +end Algebra + +variable (R S) +/-- If `R` is a domain and `S` is an arbitrary `R`-algebra, then the elements of `S` +that are algebraic over `R` form a subalgebra. -/ +def Subalgebra.algebraicClosure [IsDomain R] : Subalgebra R S where + carrier := {s | _root_.IsAlgebraic R s} + mul_mem' ha hb := ha.mul hb + add_mem' ha hb := ha.add hb + algebraMap_mem' := isAlgebraic_algebraMap + +theorem integralClosure_le_algebraicClosure [IsDomain R] : + integralClosure R S ≤ Subalgebra.algebraicClosure R S := + fun _ ↦ IsIntegral.isAlgebraic + +theorem Subalgebra.algebraicClosure_eq_integralClosure {K} [Field K] [Algebra K S] : + algebraicClosure K S = integralClosure K S := + SetLike.ext fun _ ↦ isAlgebraic_iff_isIntegral + +instance [IsDomain R] : Algebra.IsAlgebraic R (Subalgebra.algebraicClosure R S) := + (Subalgebra.isAlgebraic_iff _).mp fun _ ↦ id + +variable {R S} + +theorem Algebra.isAlgebraic_adjoin_iff [IsDomain R] {s : Set S} : + (adjoin R s).IsAlgebraic ↔ ∀ x ∈ s, IsAlgebraic R x := + Algebra.adjoin_le_iff (S := Subalgebra.algebraicClosure R S) + +theorem Algebra.isAlgebraic_adjoin_of_nonempty [NoZeroDivisors R] {s : Set S} (hs : s.Nonempty) : + (adjoin R s).IsAlgebraic ↔ ∀ x ∈ s, IsAlgebraic R x := + ⟨fun h x hx ↦ h _ (subset_adjoin hx), fun h ↦ + have ⟨x, hx⟩ := hs + have := (isDomain_iff_noZeroDivisors_and_nontrivial _).mpr ⟨‹_›, (h x hx).nontrivial⟩ + isAlgebraic_adjoin_iff.mpr h⟩ + +theorem Algebra.isAlgebraic_adjoin_singleton_iff [NoZeroDivisors R] {s : S} : + (adjoin R {s}).IsAlgebraic ↔ IsAlgebraic R s := + (isAlgebraic_adjoin_of_nonempty <| Set.singleton_nonempty s).trans forall_eq + +theorem IsAlgebraic.of_mul [NoZeroDivisors R] {y z : S} (hy : y ∈ nonZeroDivisors S) + (alg_y : IsAlgebraic R y) (alg_yz : IsAlgebraic R (y * z)) : IsAlgebraic R z := by + have ⟨t, ht, r, hr, eq⟩ := alg_y.exists_nonzero_eq_adjoin_mul hy + have := alg_yz.mul (Algebra.isAlgebraic_adjoin_singleton_iff.mpr alg_y _ ht) + rw [mul_right_comm, eq, ← Algebra.smul_def] at this + exact this.of_smul (mem_nonZeroDivisors_of_ne_zero hr) + +namespace Transcendental + +section + +variable {a : A} (ha : Transcendental R a) +include ha + +lemma extendScalars_of_isIntegral [NoZeroDivisors S] [Algebra.IsIntegral R S] + (inj : Function.Injective (algebraMap S A)) : Transcendental S a := by + contrapose ha + rw [Transcendental, not_not] at ha ⊢ + exact ha.restrictScalars_of_isIntegral _ inj + +lemma extendScalars [NoZeroDivisors S] [Algebra.IsAlgebraic R S] + (inj : Function.Injective (algebraMap S A)) : Transcendental S a := by + contrapose ha + rw [Transcendental, not_not] at ha ⊢ + exact ha.restrictScalars _ inj + +end + +variable {a : S} (ha : Transcendental R a) +include ha + +protected lemma integralClosure [NoZeroDivisors S] : + Transcendental (integralClosure R S) a := + ha.extendScalars_of_isIntegral Subtype.val_injective + +lemma subalgebraAlgebraicClosure [IsDomain R] [NoZeroDivisors S] : + Transcendental (Subalgebra.algebraicClosure R S) a := + ha.extendScalars Subtype.val_injective + +end Transcendental diff --git a/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean b/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean index b5c1195243656..22c54ebc36736 100644 --- a/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean +++ b/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean @@ -112,25 +112,11 @@ variable (A K) /-- Send a set of `x`s in a finite extension `L` of the fraction field of `R` to `(y : R) • x ∈ integralClosure R L`. -/ theorem exists_integral_multiples (s : Finset L) : - ∃ y ≠ (0 : A), ∀ x ∈ s, IsIntegral A (y • x) := by - haveI := Classical.decEq L - refine s.induction ?_ ?_ - · use 1, one_ne_zero - rintro x ⟨⟩ - · rintro x s hx ⟨y, hy, hs⟩ - have := exists_integral_multiple - ((IsFractionRing.isAlgebraic_iff A K L).mpr (.of_finite _ x)) - ((injective_iff_map_eq_zero (algebraMap A L)).mp ?_) - · rcases this with ⟨x', y', hy', hx'⟩ - refine ⟨y * y', mul_ne_zero hy hy', fun x'' hx'' => ?_⟩ - rcases Finset.mem_insert.mp hx'' with (rfl | hx'') - · rw [mul_smul, Algebra.smul_def, Algebra.smul_def, hx'] - exact isIntegral_algebraMap.mul x'.2 - · rw [mul_comm, mul_smul, Algebra.smul_def] - exact isIntegral_algebraMap.mul (hs _ hx'') - · rw [IsScalarTower.algebraMap_eq A K L] - apply (algebraMap K L).injective.comp - exact IsFractionRing.injective _ _ + ∃ y ≠ (0 : A), ∀ x ∈ s, IsIntegral A (y • x) := + have := IsLocalization.isAlgebraic K (nonZeroDivisors A) + have := Algebra.IsAlgebraic.trans' A (algebraMap K L).injective + Algebra.IsAlgebraic.exists_integral_multiples (IsScalarTower.algebraMap_eq A K L ▸ + (algebraMap K L).injective.comp (IsFractionRing.injective _ _)) _ variable (L) diff --git a/Mathlib/RingTheory/Localization/Integral.lean b/Mathlib/RingTheory/Localization/Integral.lean index 595fd2c074e5a..cb310c7e3e585 100644 --- a/Mathlib/RingTheory/Localization/Integral.lean +++ b/Mathlib/RingTheory/Localization/Integral.lean @@ -300,13 +300,11 @@ theorem isFractionRing_of_algebraic [Algebra.IsAlgebraic A L] mem_nonZeroDivisors_iff_ne_zero.mp hy ((injective_iff_map_eq_zero (algebraMap C L)).mp (algebraMap_injective C A L) _ h)) surj' := fun z => - let ⟨x, y, hy, hxy⟩ := exists_integral_multiple (Algebra.IsAlgebraic.isAlgebraic z) inj - ⟨⟨mk' C (x : L) x.2, algebraMap _ _ y, - mem_nonZeroDivisors_iff_ne_zero.mpr fun h => - hy (inj _ (by rw [IsScalarTower.algebraMap_apply A C L, h, RingHom.map_zero]))⟩, - by - simp only - rw [algebraMap_mk', ← IsScalarTower.algebraMap_apply A C L, mul_comm, hxy]⟩ + let ⟨x, hx, int⟩ := (Algebra.IsAlgebraic.isAlgebraic z).exists_integral_multiple + ((injective_iff_map_eq_zero _).mpr inj) + ⟨⟨mk' C _ int, algebraMap _ _ x, mem_nonZeroDivisors_of_ne_zero fun h ↦ + hx (inj _ <| by rw [IsScalarTower.algebraMap_apply A C L, h, RingHom.map_zero])⟩, by + rw [algebraMap_mk', ← IsScalarTower.algebraMap_apply A C L, Algebra.smul_def, mul_comm]⟩ exists_of_eq := fun {x y} h => ⟨1, by simpa using algebraMap_injective C A L h⟩ } variable (K L) @@ -320,8 +318,7 @@ theorem isFractionRing_of_finite_extension [IsDomain A] [Algebra K L] [IsScalarT isFractionRing_of_algebraic A C fun _ hx => IsFractionRing.to_map_eq_zero_iff.mp - ((_root_.map_eq_zero <| - algebraMap K L).mp <| (IsScalarTower.algebraMap_apply _ _ _ _).symm.trans hx) + ((map_eq_zero <| algebraMap K L).mp <| (IsScalarTower.algebraMap_apply _ _ _ _).symm.trans hx) end IsIntegralClosure From 66d3845e71618f19703dbe4296b90b21605c6083 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 5 Dec 2024 08:47:13 +0000 Subject: [PATCH 534/829] feat: add `HasFTaylorSeriesUpToOn.congr_series` (#19722) --- Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean b/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean index abdac0002a6e2..417b0bf00b4ec 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean @@ -146,6 +146,16 @@ theorem HasFTaylorSeriesUpToOn.congr (h : HasFTaylorSeriesUpToOn n f p s) rw [h₁ x hx] exact h.zero_eq x hx +theorem HasFTaylorSeriesUpToOn.congr_series {q} (hp : HasFTaylorSeriesUpToOn n f p s) + (hpq : ∀ m : ℕ, m ≤ n → EqOn (p · m) (q · m) s) : + HasFTaylorSeriesUpToOn n f q s where + zero_eq x hx := by simp only [← (hpq 0 (zero_le n) hx), hp.zero_eq x hx] + fderivWithin m hm x hx := by + refine ((hp.fderivWithin m hm x hx).congr' (hpq m hm.le).symm hx).congr_fderiv ?_ + refine congrArg _ (hpq (m + 1) ?_ hx) + exact ENat.add_one_natCast_le_withTop_of_lt hm + cont m hm := (hp.cont m hm).congr (hpq m hm).symm + theorem HasFTaylorSeriesUpToOn.mono (h : HasFTaylorSeriesUpToOn n f p s) {t : Set E} (hst : t ⊆ s) : HasFTaylorSeriesUpToOn n f p t := ⟨fun x hx => h.zero_eq x (hst hx), fun m hm x hx => (h.fderivWithin m hm x (hst hx)).mono hst, From 39eac4b7b275d558516e3b2c182a003b01698e02 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 5 Dec 2024 08:47:15 +0000 Subject: [PATCH 535/829] feat(FTaylorSeries): add `congr_set` lemmas (#19737) --- .../Calculus/ContDiff/FTaylorSeries.lean | 21 +++++++++++++++++++ 1 file changed, 21 insertions(+) diff --git a/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean b/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean index 417b0bf00b4ec..95ea975bb88e6 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean @@ -559,10 +559,31 @@ theorem iteratedFDerivWithin_eventually_congr_set (h : s =ᶠ[𝓝 x] t) (n : iteratedFDerivWithin 𝕜 n f s =ᶠ[𝓝 x] iteratedFDerivWithin 𝕜 n f t := iteratedFDerivWithin_eventually_congr_set' x (h.filter_mono inf_le_left) n +/-- If two sets coincide in a punctured neighborhood of `x`, +then the corresponding iterated derivatives are equal. + +Note that we also allow to puncture the neighborhood of `x` at `y`. +If `y ≠ x`, then this is a no-op. -/ +theorem iteratedFDerivWithin_congr_set' {y} (h : s =ᶠ[𝓝[{y}ᶜ] x] t) (n : ℕ) : + iteratedFDerivWithin 𝕜 n f s x = iteratedFDerivWithin 𝕜 n f t x := + (iteratedFDerivWithin_eventually_congr_set' y h n).self_of_nhds + +@[simp] +theorem iteratedFDerivWithin_insert {n y} : + iteratedFDerivWithin 𝕜 n f (insert x s) y = iteratedFDerivWithin 𝕜 n f s y := + iteratedFDerivWithin_congr_set' (y := x) + (eventually_mem_nhdsWithin.mono <| by intros; simp_all).set_eq _ + theorem iteratedFDerivWithin_congr_set (h : s =ᶠ[𝓝 x] t) (n : ℕ) : iteratedFDerivWithin 𝕜 n f s x = iteratedFDerivWithin 𝕜 n f t x := (iteratedFDerivWithin_eventually_congr_set h n).self_of_nhds +@[simp] +theorem ftaylorSeriesWithin_insert : + ftaylorSeriesWithin 𝕜 f (insert x s) = ftaylorSeriesWithin 𝕜 f s := by + ext y n : 2 + apply iteratedFDerivWithin_insert + /-- The iterated differential within a set `s` at a point `x` is not modified if one intersects `s` with a neighborhood of `x` within `s`. -/ theorem iteratedFDerivWithin_inter' {n : ℕ} (hu : u ∈ 𝓝[s] x) : From 498c8c34f7c7a2a41dd1ef4cc8295103caa07ca5 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 5 Dec 2024 09:31:14 +0000 Subject: [PATCH 536/829] refactor(Algebra/Category/ModuleCat): make `ModuleCat.Hom` a structure (#19511) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR follows the approach of #19065 to solve painful definitional equalities problems by forbidding the identification of `A ⟶ B` and `A →ₗ[R] B`. These are not equal reducibly, so tactics get confused if one is replaced by the other. There seem to be a few regressions caused by defeq abuse that we were able to get away with previously, but some subtle elaboration differences made it evident. Otherwise the cost is inserting many `.hom`s and `asHom`s. A few steps that might clean up the diff, but would be too much work to incorporate in this PR: * Make the `ext` tactic pick up `TensorProduct.ext` again. * Replace the defeq abuse between `(restrictScalars f).obj M` and `M` with explicit maps going back and forth. Overall quite a few proofs could be cleaned up at the cost of more bookkeeping. Co-authored-by: Christian Merten --- Counterexamples/Pseudoelement.lean | 22 +- .../Algebra/Category/AlgebraCat/Monoidal.lean | 7 +- .../Algebra/Category/CoalgebraCat/Basic.lean | 9 +- .../CoalgebraCat/ComonEquivalence.lean | 36 +- .../Category/CoalgebraCat/Monoidal.lean | 6 +- .../Algebra/Category/FGModuleCat/Basic.lean | 65 ++- .../Algebra/Category/FGModuleCat/Limits.lean | 4 +- Mathlib/Algebra/Category/Grp/Injective.lean | 4 +- .../Category/Grp/ZModuleEquivalence.lean | 13 +- .../Algebra/Category/ModuleCat/Abelian.lean | 20 +- .../Category/ModuleCat/Adjunctions.lean | 19 +- .../Algebra/Category/ModuleCat/Algebra.lean | 13 +- Mathlib/Algebra/Category/ModuleCat/Basic.lean | 408 ++++++++++++------ .../Category/ModuleCat/Biproducts.lean | 42 +- .../Category/ModuleCat/ChangeOfRings.lean | 393 ++++++++--------- .../ModuleCat/Differentials/Basic.lean | 10 +- .../ModuleCat/Differentials/Presheaf.lean | 4 +- .../Algebra/Category/ModuleCat/EpiMono.lean | 20 +- .../Category/ModuleCat/FilteredColimits.lean | 17 +- Mathlib/Algebra/Category/ModuleCat/Free.lean | 12 +- .../Algebra/Category/ModuleCat/Images.lean | 45 +- .../Algebra/Category/ModuleCat/Injective.lean | 9 +- .../Algebra/Category/ModuleCat/Kernels.lean | 42 +- .../Algebra/Category/ModuleCat/Limits.lean | 38 +- .../Category/ModuleCat/Monoidal/Basic.lean | 114 ++--- .../Category/ModuleCat/Monoidal/Closed.lean | 28 +- .../ModuleCat/Monoidal/Symmetric.lean | 8 +- .../Algebra/Category/ModuleCat/Presheaf.lean | 71 +-- .../ModuleCat/Presheaf/ChangeOfRings.lean | 7 +- .../Category/ModuleCat/Presheaf/Monoidal.lean | 10 +- .../ModuleCat/Presheaf/Pushforward.lean | 35 +- .../Category/ModuleCat/Presheaf/Sheafify.lean | 12 +- .../Algebra/Category/ModuleCat/Products.lean | 25 +- .../Category/ModuleCat/Projective.lean | 19 +- Mathlib/Algebra/Category/ModuleCat/Sheaf.lean | 2 +- .../ModuleCat/Sheaf/ChangeOfRings.lean | 2 +- .../Algebra/Category/ModuleCat/Subobject.lean | 52 +-- .../Algebra/Category/ModuleCat/Tannaka.lean | 15 +- Mathlib/Algebra/Homology/LocalCohomology.lean | 4 +- .../Homology/ShortComplex/ModuleCat.lean | 69 +-- Mathlib/AlgebraicGeometry/AffineSpace.lean | 4 +- Mathlib/AlgebraicGeometry/Modules/Tilde.lean | 73 ++-- .../Abelian/Pseudoelements.lean | 12 +- .../WithAlgebraicStructures.lean | 46 +- Mathlib/CategoryTheory/Linear/Yoneda.lean | 3 - .../Monoidal/Internal/Module.lean | 82 ++-- Mathlib/Condensed/Discrete/Module.lean | 12 +- .../QuadraticForm/QuadraticModuleCat.lean | 10 +- Mathlib/RepresentationTheory/Character.lean | 4 +- Mathlib/RepresentationTheory/FDRep.lean | 28 +- .../GroupCohomology/Basic.lean | 34 +- .../GroupCohomology/LowDegree.lean | 34 +- .../GroupCohomology/Resolution.lean | 31 +- Mathlib/RepresentationTheory/Invariants.lean | 14 +- Mathlib/RepresentationTheory/Rep.lean | 125 +++--- .../AdicCompletion/AsTensorProduct.lean | 16 +- .../RingTheory/Coalgebra/TensorProduct.lean | 4 +- Mathlib/RingTheory/Flat/CategoryTheory.lean | 4 +- .../Topology/Category/Profinite/Nobeling.lean | 6 +- .../ConcreteCategory/ModuleCat.lean | 53 +++ 60 files changed, 1291 insertions(+), 1035 deletions(-) create mode 100644 MathlibTest/CategoryTheory/ConcreteCategory/ModuleCat.lean diff --git a/Counterexamples/Pseudoelement.lean b/Counterexamples/Pseudoelement.lean index af58937a6308c..46052cc97e77a 100644 --- a/Counterexamples/Pseudoelement.lean +++ b/Counterexamples/Pseudoelement.lean @@ -74,26 +74,24 @@ theorem x_not_pseudo_eq : ¬PseudoEqual _ x y := by replace h := ModuleCat.eq_range_of_pseudoequal h dsimp [x, y] at h let φ := biprod.lift (𝟙 (of ℤ ℚ)) (2 • 𝟙 (of ℤ ℚ)) - have mem_range := mem_range_self φ (1 : ℚ) + have mem_range := mem_range_self φ.hom (1 : ℚ) rw [h] at mem_range obtain ⟨a, ha⟩ := mem_range - erw [← ModuleCat.id_apply (φ (1 : ℚ)), ← biprod.total, ← LinearMap.comp_apply, ← comp_def, - Preadditive.comp_add] at ha + rw [← ModuleCat.id_apply _ (φ (1 : ℚ)), ← biprod.total, ← LinearMap.comp_apply, + ← ModuleCat.hom_comp, Preadditive.comp_add] at ha let π₁ := (biprod.fst : of ℤ ℚ ⊞ of ℤ ℚ ⟶ _) have ha₁ := congr_arg π₁ ha - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [← CategoryTheory.comp_apply, ← CategoryTheory.comp_apply] at ha₁ + rw [← ModuleCat.comp_apply, ← ModuleCat.comp_apply] at ha₁ simp only [π₁, φ, BinaryBiproduct.bicone_fst, biprod.lift_fst, CategoryTheory.id_apply, biprod.lift_fst_assoc, Category.id_comp, biprod.lift_snd_assoc, Linear.smul_comp, Preadditive.add_comp, BinaryBicone.inl_fst, BinaryBicone.inr_fst, smul_zero, add_zero] at ha₁ let π₂ := (biprod.snd : of ℤ ℚ ⊞ of ℤ ℚ ⟶ _) have ha₂ := congr_arg π₂ ha - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [← CategoryTheory.comp_apply, ← CategoryTheory.comp_apply] at ha₂ + rw [← ModuleCat.comp_apply, ← ModuleCat.comp_apply] at ha₂ simp only [π₁, π₂, φ, BinaryBiproduct.bicone_snd, biprod.lift_snd, CategoryTheory.id_apply, biprod.lift_fst_assoc, Category.id_comp, biprod.lift_snd_assoc, Linear.smul_comp, Preadditive.add_comp, BinaryBicone.inl_snd, BinaryBicone.inr_snd, zero_add, two_smul] at ha₂ - erw [add_apply, CategoryTheory.id_apply] at ha₂ + erw [add_apply, ModuleCat.id_apply] at ha₂ subst ha₁ simp only [self_eq_add_right] at ha₂ exact one_ne_zero' ℚ ha₂ @@ -104,12 +102,12 @@ open scoped Pseudoelement /-- `biprod.fst ⟦x⟧ = biprod.fst ⟦y⟧`. -/ theorem fst_mk'_x_eq_fst_mk'_y : - (biprod.fst : of ℤ ℚ ⊞ of ℤ ℚ ⟶ _) ⟦x⟧ = (biprod.fst : of ℤ ℚ ⊞ of ℤ ℚ ⟶ _) ⟦y⟧ := + pseudoApply biprod.fst ⟦x⟧ = pseudoApply biprod.fst ⟦y⟧ := Quotient.eq.2 fst_x_pseudo_eq_fst_y /-- `biprod.snd ⟦x⟧ = biprod.snd ⟦y⟧`. -/ theorem snd_mk'_x_eq_snd_mk'_y : - (biprod.snd : of ℤ ℚ ⊞ of ℤ ℚ ⟶ _) ⟦x⟧ = (biprod.snd : of ℤ ℚ ⊞ of ℤ ℚ ⟶ _) ⟦y⟧ := + pseudoApply biprod.snd ⟦x⟧ = pseudoApply biprod.snd ⟦y⟧ := Quotient.eq.2 snd_x_pseudo_eq_snd_y -- Porting note: needs explicit type ascription `: Quotient <| Pseudoelement.setoid _` @@ -125,8 +123,8 @@ theorem exist_ne_and_fst_eq_fst_and_snd_eq_snd : ∃ x y, -- Porting note: removed type ascription `: of ℤ ℚ ⊞ of ℤ ℚ`, it gave an error about -- `Type` not having zero morphisms. jmc: I don't understand where the error came from x ≠ y ∧ - (biprod.fst : of ℤ ℚ ⊞ of ℤ ℚ ⟶ _) x = (biprod.fst : of ℤ ℚ ⊞ of ℤ ℚ ⟶ _) y ∧ - (biprod.snd : of ℤ ℚ ⊞ of ℤ ℚ ⟶ _) x = (biprod.snd : of ℤ ℚ ⊞ of ℤ ℚ ⟶ _) y := + pseudoApply (biprod.fst : of ℤ ℚ ⊞ of ℤ ℚ ⟶ _) x = pseudoApply biprod.fst y ∧ + pseudoApply biprod.snd x = pseudoApply biprod.snd y := ⟨⟦x⟧, ⟦y⟧, mk'_x_ne_mk'_y, fst_mk'_x_eq_fst_mk'_y, snd_mk'_x_eq_snd_mk'_y⟩ end diff --git a/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean b/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean index fa16a8aac4bcd..48a1d4c9ff2f5 100644 --- a/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean +++ b/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean @@ -60,9 +60,10 @@ noncomputable instance instMonoidalCategory : MonoidalCategory (AlgebraCat.{u} R (forget₂ (AlgebraCat R) (ModuleCat R)) { μIso := fun _ _ => Iso.refl _ εIso := Iso.refl _ - associator_eq := fun _ _ _ => TensorProduct.ext_threefold (fun _ _ _ => rfl) - leftUnitor_eq := fun _ => TensorProduct.ext' (fun _ _ => rfl) - rightUnitor_eq := fun _ => TensorProduct.ext' (fun _ _ => rfl) } + associator_eq := fun _ _ _ => + ModuleCat.hom_ext <| TensorProduct.ext_threefold (fun _ _ _ => rfl) + leftUnitor_eq := fun _ => ModuleCat.hom_ext <| TensorProduct.ext' (fun _ _ => rfl) + rightUnitor_eq := fun _ => ModuleCat.hom_ext <| TensorProduct.ext' (fun _ _ => rfl) } /-- `forget₂ (AlgebraCat R) (ModuleCat R)` as a monoidal functor. -/ example : (forget₂ (AlgebraCat R) (ModuleCat R)).Monoidal := inferInstance diff --git a/Mathlib/Algebra/Category/CoalgebraCat/Basic.lean b/Mathlib/Algebra/Category/CoalgebraCat/Basic.lean index 3d24b8db74884..362312b348546 100644 --- a/Mathlib/Algebra/Category/CoalgebraCat/Basic.lean +++ b/Mathlib/Algebra/Category/CoalgebraCat/Basic.lean @@ -46,8 +46,9 @@ variable (R) /-- The object in the category of `R`-coalgebras associated to an `R`-coalgebra. -/ @[simps] def of (X : Type v) [AddCommGroup X] [Module R X] [Coalgebra R X] : - CoalgebraCat R where - instCoalgebra := (inferInstance : Coalgebra R X) + CoalgebraCat R := + { ModuleCat.of R X with + instCoalgebra := (inferInstance : Coalgebra R X) } variable {R} @@ -105,7 +106,7 @@ instance concreteCategory : ConcreteCategory.{v} (CoalgebraCat.{v} R) where instance hasForgetToModule : HasForget₂ (CoalgebraCat R) (ModuleCat R) where forget₂ := { obj := fun M => ModuleCat.of R M - map := fun f => f.toCoalgHom.toLinearMap } + map := fun f => ModuleCat.ofHom f.toCoalgHom.toLinearMap } @[simp] theorem forget₂_obj (X : CoalgebraCat R) : @@ -114,7 +115,7 @@ theorem forget₂_obj (X : CoalgebraCat R) : @[simp] theorem forget₂_map (X Y : CoalgebraCat R) (f : X ⟶ Y) : - (forget₂ (CoalgebraCat R) (ModuleCat R)).map f = (f.toCoalgHom : X →ₗ[R] Y) := + (forget₂ (CoalgebraCat R) (ModuleCat R)).map f = ModuleCat.ofHom (f.toCoalgHom : X →ₗ[R] Y) := rfl end CoalgebraCat diff --git a/Mathlib/Algebra/Category/CoalgebraCat/ComonEquivalence.lean b/Mathlib/Algebra/Category/CoalgebraCat/ComonEquivalence.lean index 322fce71e773d..51ff35c7035db 100644 --- a/Mathlib/Algebra/Category/CoalgebraCat/ComonEquivalence.lean +++ b/Mathlib/Algebra/Category/CoalgebraCat/ComonEquivalence.lean @@ -36,13 +36,13 @@ open CategoryTheory MonoidalCategory variable {R : Type u} [CommRing R] /-- An `R`-coalgebra is a comonoid object in the category of `R`-modules. -/ -@[simps] def toComonObj (X : CoalgebraCat R) : Comon_ (ModuleCat R) where +@[simps X counit comul] def toComonObj (X : CoalgebraCat R) : Comon_ (ModuleCat R) where X := ModuleCat.of R X counit := ModuleCat.ofHom Coalgebra.counit comul := ModuleCat.ofHom Coalgebra.comul - counit_comul := by simpa only [ModuleCat.of_coe] using Coalgebra.rTensor_counit_comp_comul - comul_counit := by simpa only [ModuleCat.of_coe] using Coalgebra.lTensor_counit_comp_comul - comul_assoc := by simp_rw [ModuleCat.of_coe]; exact Coalgebra.coassoc.symm + counit_comul := ModuleCat.hom_ext <| by simpa using Coalgebra.rTensor_counit_comp_comul + comul_counit := ModuleCat.hom_ext <| by simpa using Coalgebra.lTensor_counit_comp_comul + comul_assoc := ModuleCat.hom_ext <| by simp_rw [ModuleCat.of_coe]; exact Coalgebra.coassoc.symm variable (R) in /-- The natural functor from `R`-coalgebras to comonoid objects in the category of `R`-modules. -/ @@ -51,26 +51,26 @@ def toComon : CoalgebraCat R ⥤ Comon_ (ModuleCat R) where obj X := toComonObj X map f := { hom := ModuleCat.ofHom f.1 - hom_counit := f.1.counit_comp - hom_comul := f.1.map_comp_comul.symm } + hom_counit := ModuleCat.hom_ext f.1.counit_comp + hom_comul := ModuleCat.hom_ext f.1.map_comp_comul.symm } /-- A comonoid object in the category of `R`-modules has a natural comultiplication and counit. -/ @[simps] instance ofComonObjCoalgebraStruct (X : Comon_ (ModuleCat R)) : CoalgebraStruct R X.X where - comul := X.comul - counit := X.counit + comul := X.comul.hom + counit := X.counit.hom /-- A comonoid object in the category of `R`-modules has a natural `R`-coalgebra structure. -/ -def ofComonObj (X : Comon_ (ModuleCat R)) : CoalgebraCat R where - carrier := X.X - instCoalgebra := - { ofComonObjCoalgebraStruct X with - coassoc := X.comul_assoc.symm - rTensor_counit_comp_comul := X.counit_comul - lTensor_counit_comp_comul := X.comul_counit } +def ofComonObj (X : Comon_ (ModuleCat R)) : CoalgebraCat R := + { ModuleCat.of R X.X with + instCoalgebra := + { ofComonObjCoalgebraStruct X with + coassoc := ModuleCat.hom_ext_iff.mp X.comul_assoc.symm + rTensor_counit_comp_comul := ModuleCat.hom_ext_iff.mp X.counit_comul + lTensor_counit_comp_comul := ModuleCat.hom_ext_iff.mp X.comul_counit } } variable (R) @@ -79,9 +79,9 @@ def ofComon : Comon_ (ModuleCat R) ⥤ CoalgebraCat R where obj X := ofComonObj X map f := { toCoalgHom := - { f.hom with - counit_comp := f.hom_counit - map_comp_comul := f.hom_comul.symm }} + { f.hom.hom with + counit_comp := ModuleCat.hom_ext_iff.mp f.hom_counit + map_comp_comul := ModuleCat.hom_ext_iff.mp f.hom_comul.symm }} /-- The natural category equivalence between `R`-coalgebras and comonoid objects in the category of `R`-modules. -/ diff --git a/Mathlib/Algebra/Category/CoalgebraCat/Monoidal.lean b/Mathlib/Algebra/Category/CoalgebraCat/Monoidal.lean index dbd37c50e0b59..be3ac68381c8e 100644 --- a/Mathlib/Algebra/Category/CoalgebraCat/Monoidal.lean +++ b/Mathlib/Algebra/Category/CoalgebraCat/Monoidal.lean @@ -50,9 +50,9 @@ noncomputable def MonoidalCategory.inducingFunctorData : whiskerRight_eq X f := by ext; rfl tensorHom_eq f g := by ext; rfl εIso := Iso.refl _ - associator_eq X Y Z := TensorProduct.ext <| TensorProduct.ext <| by ext; rfl - leftUnitor_eq X := TensorProduct.ext <| by ext; rfl - rightUnitor_eq X := TensorProduct.ext <| by ext; rfl + associator_eq X Y Z := ModuleCat.hom_ext <| TensorProduct.ext <| TensorProduct.ext <| by ext; rfl + leftUnitor_eq X := ModuleCat.hom_ext <| TensorProduct.ext <| by ext; rfl + rightUnitor_eq X := ModuleCat.hom_ext <| TensorProduct.ext <| by ext; rfl noncomputable instance instMonoidalCategory : MonoidalCategory (CoalgebraCat R) := Monoidal.induced (forget₂ _ (ModuleCat R)) (MonoidalCategory.inducingFunctorData R) diff --git a/Mathlib/Algebra/Category/FGModuleCat/Basic.lean b/Mathlib/Algebra/Category/FGModuleCat/Basic.lean index 613eb1a083c3f..07f1ca89ad5e6 100644 --- a/Mathlib/Algebra/Category/FGModuleCat/Basic.lean +++ b/Mathlib/Algebra/Category/FGModuleCat/Basic.lean @@ -72,12 +72,6 @@ instance : LargeCategory (FGModuleCat R) := by dsimp [FGModuleCat] infer_instance -instance {M N : FGModuleCat R} : FunLike (M ⟶ N) M N := - LinearMap.instFunLike - -instance {M N : FGModuleCat R} : LinearMapClass (M ⟶ N) R M N := - LinearMap.semilinearMapClass - instance : ConcreteCategory (FGModuleCat R) := by dsimp [FGModuleCat] infer_instance @@ -101,9 +95,20 @@ instance : Inhabited (FGModuleCat R) := ⟨⟨ModuleCat.of R R, Module.Finite.self R⟩⟩ /-- Lift an unbundled finitely generated module to `FGModuleCat R`. -/ -def of (V : Type u) [AddCommGroup V] [Module R V] [Module.Finite R V] : FGModuleCat R := +abbrev of (V : Type u) [AddCommGroup V] [Module R V] [Module.Finite R V] : FGModuleCat R := ⟨ModuleCat.of R V, by change Module.Finite R V; infer_instance⟩ +variable {R} in +/-- Lift a linear map between finitely generated modules to `FGModuleCat R`. -/ +abbrev ofHom {V W : Type u} [AddCommGroup V] [Module R V] [Module.Finite R V] + [AddCommGroup W] [Module R W] [Module.Finite R W] + (f : V →ₗ[R] W) : of R V ⟶ of R W := + ModuleCat.ofHom f + +variable {R} in +@[ext] lemma hom_ext {V W : FGModuleCat R} {f g : V ⟶ W} (h : f.hom = g.hom) : f = g := + ModuleCat.hom_ext h + instance (V : FGModuleCat R) : Module.Finite R V := V.property @@ -127,8 +132,8 @@ def _root_.LinearEquiv.toFGModuleCatIso {V W : Type u} [AddCommGroup V] [Module R V] [Module.Finite R V] [AddCommGroup W] [Module R W] [Module.Finite R W] (e : V ≃ₗ[R] W) : FGModuleCat.of R V ≅ FGModuleCat.of R W where - hom := e.toLinearMap - inv := e.symm.toLinearMap + hom := ModuleCat.ofHom e.toLinearMap + inv := ModuleCat.ofHom e.symm.toLinearMap hom_inv_id := by ext x; exact e.left_inv x inv_hom_id := by ext x; exact e.right_inv x @@ -176,7 +181,11 @@ instance : (forget₂ (FGModuleCat.{u} R) (ModuleCat.{u} R)).Additive where instance : (forget₂ (FGModuleCat.{u} R) (ModuleCat.{u} R)).Linear R where theorem Iso.conj_eq_conj {V W : FGModuleCat R} (i : V ≅ W) (f : End V) : - Iso.conj i f = LinearEquiv.conj (isoToLinearEquiv i) f := + Iso.conj i f = FGModuleCat.ofHom (LinearEquiv.conj (isoToLinearEquiv i) f.hom) := + rfl + +theorem Iso.conj_hom_eq_conj {V W : FGModuleCat R} (i : V ≅ W) (f : End V) : + (Iso.conj i f).hom = (LinearEquiv.conj (isoToLinearEquiv i) f.hom) := rfl end CommRing @@ -186,11 +195,12 @@ section Field variable (K : Type u) [Field K] instance (V W : FGModuleCat K) : Module.Finite K (V ⟶ W) := - (by infer_instance : Module.Finite K (V →ₗ[K] W)) + (inferInstanceAs <| Module.Finite K (V →ₗ[K] W)).equiv ModuleCat.homLinearEquiv.symm instance closedPredicateModuleFinite : MonoidalCategory.ClosedPredicate fun V : ModuleCat.{u} K ↦ Module.Finite K V where - prop_ihom {X Y} _ _ := Module.Finite.linearMap K K X Y + prop_ihom {X Y} _ _ := + (inferInstanceAs <| Module.Finite K (X →ₗ[K] Y)).equiv ModuleCat.homLinearEquiv.symm instance : MonoidalClosed (FGModuleCat K) := by dsimp [FGModuleCat] @@ -201,7 +211,7 @@ instance : MonoidalClosed (FGModuleCat K) := by variable (V W : FGModuleCat K) @[simp] -theorem ihom_obj : (ihom V).obj W = FGModuleCat.of K (V →ₗ[K] W) := +theorem ihom_obj : (ihom V).obj W = FGModuleCat.of K (V ⟶ W) := rfl /-- The dual module is the dual in the rigid monoidal category `FGModuleCat K`. -/ @@ -216,33 +226,43 @@ open CategoryTheory.MonoidalCategory /-- The coevaluation map is defined in `LinearAlgebra.coevaluation`. -/ def FGModuleCatCoevaluation : 𝟙_ (FGModuleCat K) ⟶ V ⊗ FGModuleCatDual K V := - coevaluation K V + ModuleCat.ofHom <| coevaluation K V theorem FGModuleCatCoevaluation_apply_one : - FGModuleCatCoevaluation K V (1 : K) = + (FGModuleCatCoevaluation K V).hom (1 : K) = ∑ i : Basis.ofVectorSpaceIndex K V, (Basis.ofVectorSpace K V) i ⊗ₜ[K] (Basis.ofVectorSpace K V).coord i := coevaluation_apply_one K V /-- The evaluation morphism is given by the contraction map. -/ def FGModuleCatEvaluation : FGModuleCatDual K V ⊗ V ⟶ 𝟙_ (FGModuleCat K) := - contractLeft K V + ModuleCat.ofHom <| contractLeft K V -@[simp] theorem FGModuleCatEvaluation_apply (f : FGModuleCatDual K V) (x : V) : - (FGModuleCatEvaluation K V) (f ⊗ₜ x) = f.toFun x := + (FGModuleCatEvaluation K V).hom (f ⊗ₜ x) = f.toFun x := + contractLeft_apply f x + +/-- `@[simp]`-normal form of `FGModuleCatEvaluation_apply`, where the carriers have been unfolded. +-/ +@[simp] +theorem FGModuleCatEvaluation_apply' (f : FGModuleCatDual K V) (x : V) : + DFunLike.coe + (F := ((ModuleCat.of K (Module.Dual K V) ⊗ V.obj).carrier →ₗ[K] (𝟙_ (ModuleCat K)))) + (FGModuleCatEvaluation K V).hom (f ⊗ₜ x) = f.toFun x := contractLeft_apply f x private theorem coevaluation_evaluation : letI V' : FGModuleCat K := FGModuleCatDual K V V' ◁ FGModuleCatCoevaluation K V ≫ (α_ V' V V').inv ≫ FGModuleCatEvaluation K V ▷ V' = (ρ_ V').hom ≫ (λ_ V').inv := by + ext : 1 apply contractLeft_assoc_coevaluation K V private theorem evaluation_coevaluation : FGModuleCatCoevaluation K V ▷ V ≫ (α_ V (FGModuleCatDual K V) V).hom ≫ V ◁ FGModuleCatEvaluation K V = (λ_ V).hom ≫ (ρ_ V).inv := by + ext : 1 apply contractLeft_assoc_coevaluation' K V instance exactPairing : ExactPairing V (FGModuleCatDual K V) where @@ -266,9 +286,10 @@ end FGModuleCat @[simp] theorem LinearMap.comp_id_fgModuleCat {R} [Ring R] {G : FGModuleCat.{u} R} {H : Type u} [AddCommGroup H] [Module R H] - (f : G →ₗ[R] H) : f.comp (𝟙 G) = f := - Category.id_comp (ModuleCat.ofHom f) + (f : G →ₗ[R] H) : f.comp (ModuleCat.Hom.hom (𝟙 G)) = f := + ModuleCat.hom_ext_iff.mp <| Category.id_comp (ModuleCat.ofHom f) + @[simp] theorem LinearMap.id_fgModuleCat_comp {R} [Ring R] {G : Type u} [AddCommGroup G] [Module R G] {H : FGModuleCat.{u} R} - (f : G →ₗ[R] H) : LinearMap.comp (𝟙 H) f = f := - Category.comp_id (ModuleCat.ofHom f) + (f : G →ₗ[R] H) : LinearMap.comp (ModuleCat.Hom.hom (𝟙 H)) f = f := + ModuleCat.hom_ext_iff.mp <| Category.comp_id (ModuleCat.ofHom f) diff --git a/Mathlib/Algebra/Category/FGModuleCat/Limits.lean b/Mathlib/Algebra/Category/FGModuleCat/Limits.lean index 3e5af1918343b..7b88d667c1bdb 100644 --- a/Mathlib/Algebra/Category/FGModuleCat/Limits.lean +++ b/Mathlib/Algebra/Category/FGModuleCat/Limits.lean @@ -38,7 +38,7 @@ variable {k : Type v} [Field k] instance {J : Type} [Finite J] (Z : J → ModuleCat.{v} k) [∀ j, FiniteDimensional k (Z j)] : FiniteDimensional k (∏ᶜ fun j => Z j : ModuleCat.{v} k) := haveI : FiniteDimensional k (ModuleCat.of k (∀ j, Z j)) := by unfold ModuleCat.of; infer_instance - FiniteDimensional.of_injective (ModuleCat.piIsoPi _).hom + FiniteDimensional.of_injective (ModuleCat.piIsoPi _).hom.hom ((ModuleCat.mono_iff_injective _).1 (by infer_instance)) /-- Finite limits of finite dimensional vectors spaces are finite dimensional, @@ -48,7 +48,7 @@ instance (F : J ⥤ FGModuleCat k) : haveI : ∀ j, FiniteDimensional k ((F ⋙ forget₂ (FGModuleCat k) (ModuleCat.{v} k)).obj j) := by intro j; change FiniteDimensional k (F.obj j); infer_instance FiniteDimensional.of_injective - (limitSubobjectProduct (F ⋙ forget₂ (FGModuleCat k) (ModuleCat.{v} k))) + (limitSubobjectProduct (F ⋙ forget₂ (FGModuleCat k) (ModuleCat.{v} k))).hom ((ModuleCat.mono_iff_injective _).1 inferInstance) /-- The forgetful functor from `FGModuleCat k` to `ModuleCat k` creates all finite limits. -/ diff --git a/Mathlib/Algebra/Category/Grp/Injective.lean b/Mathlib/Algebra/Category/Grp/Injective.lean index f8ea96a6004ef..a5f8e576a886a 100644 --- a/Mathlib/Algebra/Category/Grp/Injective.lean +++ b/Mathlib/Algebra/Category/Grp/Injective.lean @@ -57,9 +57,9 @@ theorem Module.Baer.of_divisible [DivisibleBy A ℤ] : Module.Baer ℤ A := fun namespace AddCommGrp -theorem injective_as_module_iff : Injective (⟨A⟩ : ModuleCat ℤ) ↔ +theorem injective_as_module_iff : Injective (ModuleCat.of ℤ A) ↔ Injective (⟨A,inferInstance⟩ : AddCommGrp) := - ((forget₂ (ModuleCat ℤ) AddCommGrp).asEquivalence.map_injective_iff ⟨A⟩).symm + ((forget₂ (ModuleCat ℤ) AddCommGrp).asEquivalence.map_injective_iff (ModuleCat.of ℤ A)).symm instance injective_of_divisible [DivisibleBy A ℤ] : Injective (⟨A,inferInstance⟩ : AddCommGrp) := diff --git a/Mathlib/Algebra/Category/Grp/ZModuleEquivalence.lean b/Mathlib/Algebra/Category/Grp/ZModuleEquivalence.lean index 242349db0fa0c..cd88f753ea5b8 100644 --- a/Mathlib/Algebra/Category/Grp/ZModuleEquivalence.lean +++ b/Mathlib/Algebra/Category/Grp/ZModuleEquivalence.lean @@ -29,12 +29,13 @@ instance forget₂_addCommGroup_full : (forget₂ (ModuleCat ℤ) AddCommGrp.{u} -- `AddMonoidHom.toIntLinearMap` doesn't work here because `A` and `B` are not -- definitionally equal to the canonical `AddCommGroup.toIntModule` module -- instances it expects. - f := ⟨@LinearMap.mk _ _ _ _ _ _ _ _ _ A.isModule B.isModule - { toFun := f, - map_add' := AddMonoidHom.map_add (show A.carrier →+ B.carrier from f) } - (fun n x => by - convert AddMonoidHom.map_zsmul (show A.carrier →+ B.carrier from f) x n <;> - ext <;> apply int_smul_eq_zsmul), rfl⟩ + f := ⟨@ModuleCat.ofHom _ _ _ _ _ A.isModule _ B.isModule <| + @LinearMap.mk _ _ _ _ _ _ _ _ _ A.isModule B.isModule + { toFun := f, + map_add' := AddMonoidHom.map_add (show A.carrier →+ B.carrier from f) } + (fun n x => by + convert AddMonoidHom.map_zsmul (show A.carrier →+ B.carrier from f) x n <;> + ext <;> apply int_smul_eq_zsmul), rfl⟩ /-- The forgetful functor from `ℤ` modules to `AddCommGrp` is essentially surjective. -/ instance forget₂_addCommGrp_essSurj : (forget₂ (ModuleCat ℤ) AddCommGrp.{u}).EssSurj where diff --git a/Mathlib/Algebra/Category/ModuleCat/Abelian.lean b/Mathlib/Algebra/Category/ModuleCat/Abelian.lean index fcb99e727a6c4..2bb2a6d564458 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Abelian.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Abelian.lean @@ -29,9 +29,9 @@ variable {R : Type u} [Ring R] {M N : ModuleCat.{v} R} (f : M ⟶ N) /-- In the category of modules, every monomorphism is normal. -/ def normalMono (hf : Mono f) : NormalMono f where - Z := of R (N ⧸ LinearMap.range f) - g := f.range.mkQ - w := LinearMap.range_mkQ_comp _ + Z := of R (N ⧸ LinearMap.range f.hom) + g := ofHom f.hom.range.mkQ + w := hom_ext <| LinearMap.range_mkQ_comp _ isLimit := /- The following [invalid Lean code](https://github.com/leanprover-community/lean/issues/341) might help you understand what's going on here: @@ -43,16 +43,16 @@ def normalMono (hf : Mono f) : NormalMono f where ``` -/ IsKernel.isoKernel _ _ (kernelIsLimit _) - (LinearEquiv.toModuleIso' + (LinearEquiv.toModuleIso ((Submodule.quotEquivOfEqBot _ (ker_eq_bot_of_mono _)).symm ≪≫ₗ - (LinearMap.quotKerEquivRange f ≪≫ₗ + (LinearMap.quotKerEquivRange f.hom ≪≫ₗ LinearEquiv.ofEq _ _ (Submodule.ker_mkQ _).symm))) <| by ext; rfl /-- In the category of modules, every epimorphism is normal. -/ def normalEpi (hf : Epi f) : NormalEpi f where - W := of R (LinearMap.ker f) - g := (LinearMap.ker f).subtype - w := LinearMap.comp_ker_subtype _ + W := of R (LinearMap.ker f.hom) + g := ofHom (LinearMap.ker f.hom).subtype + w := hom_ext <| LinearMap.comp_ker_subtype _ isColimit := /- The following invalid Lean code might help you understand what's going on here: ``` @@ -63,9 +63,9 @@ def normalEpi (hf : Epi f) : NormalEpi f where ``` -/ IsCokernel.cokernelIso _ _ (cokernelIsColimit _) - (LinearEquiv.toModuleIso' + (LinearEquiv.toModuleIso (Submodule.quotEquivOfEq _ _ (Submodule.range_subtype _) ≪≫ₗ - LinearMap.quotKerEquivRange f ≪≫ₗ + LinearMap.quotKerEquivRange f.hom ≪≫ₗ LinearEquiv.ofTop _ (range_eq_top_of_epi _))) <| by ext; rfl /-- The category of R-modules is abelian. -/ diff --git a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean index 12126e451178b..41760fe6fdefd 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean @@ -36,9 +36,9 @@ free `R`-module with generators `x : X`, implemented as the type `X →₀ R`. -/ def free : Type u ⥤ ModuleCat R where obj X := ModuleCat.of R (X →₀ R) - map {_ _} f := Finsupp.lmapDomain _ _ f - map_id := by intros; exact Finsupp.lmapDomain_id _ _ - map_comp := by intros; exact Finsupp.lmapDomain_comp _ _ _ _ + map {_ _} f := ofHom <| Finsupp.lmapDomain _ _ f + map_id := by intros; ext : 1; exact Finsupp.lmapDomain_id _ _ + map_comp := by intros; ext : 1; exact Finsupp.lmapDomain_comp _ _ _ _ variable {R} @@ -49,13 +49,13 @@ noncomputable def freeMk {X : Type u} (x : X) : (free R).obj X := Finsupp.single lemma free_hom_ext {X : Type u} {M : ModuleCat.{u} R} {f g : (free R).obj X ⟶ M} (h : ∀ (x : X), f (freeMk x) = g (freeMk x)) : f = g := - (Finsupp.lhom_ext' (fun x ↦ LinearMap.ext_ring (h x))) + ModuleCat.hom_ext (Finsupp.lhom_ext' (fun x ↦ LinearMap.ext_ring (h x))) /-- The morphism of modules `(free R).obj X ⟶ M` corresponding to a map `f : X ⟶ M`. -/ noncomputable def freeDesc {X : Type u} {M : ModuleCat.{u} R} (f : X ⟶ M) : (free R).obj X ⟶ M := - Finsupp.lift M R X f + ofHom <| Finsupp.lift M R X f @[simp] lemma freeDesc_apply {X : Type u} {M : ModuleCat.{u} R} (f : X ⟶ M) (x : X) : @@ -109,12 +109,11 @@ namespace FreeMonoidal (This should not be used directly: it is part of the implementation of the monoidal structure on the functor `free R`.) -/ def εIso : 𝟙_ (ModuleCat R) ≅ (free R).obj (𝟙_ (Type u)) where - hom := Finsupp.lsingle PUnit.unit - inv := Finsupp.lapply PUnit.unit + hom := ofHom <| Finsupp.lsingle PUnit.unit + inv := ofHom <| Finsupp.lapply PUnit.unit hom_inv_id := by - ext x - dsimp - erw [Finsupp.lapply_apply, Finsupp.lsingle_apply, Finsupp.single_eq_same] + ext + simp [free] inv_hom_id := by ext ⟨⟩ dsimp [freeMk] diff --git a/Mathlib/Algebra/Category/ModuleCat/Algebra.lean b/Mathlib/Algebra/Category/ModuleCat/Algebra.lean index d2d388822bfff..275ae2fa12b4f 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Algebra.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Algebra.lean @@ -51,18 +51,9 @@ theorem isScalarTower_of_algebra_moduleCat (M : ModuleCat.{v} A) : IsScalarTower attribute [scoped instance] ModuleCat.isScalarTower_of_algebra_moduleCat -- We verify that the morphism spaces become `k`-modules. -example (M N : ModuleCat.{v} A) : Module k (M ⟶ N) := LinearMap.module --- Porting note: used to be `by infer_instance` instead of `LinearMap.module` +example (M N : ModuleCat.{v} A) : Module k (M ⟶ N) := inferInstance instance linearOverField : Linear k (ModuleCat.{v} A) where - -- Porting note: used to be `by infer_instance` instead of `LinearMap.module` - homModule _ _ := LinearMap.module - smul_comp := by - -- Porting note: this was automatic by `aesop_cat` - intros - ext - dsimp only [coe_comp, Function.comp_apply] - rw [LinearMap.smul_apply, LinearMap.map_smul_of_tower] - rfl + homModule _ _ := inferInstance end ModuleCat diff --git a/Mathlib/Algebra/Category/ModuleCat/Basic.lean b/Mathlib/Algebra/Category/ModuleCat/Basic.lean index 5a22d8f540260..9d0fdd30373c1 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Basic.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Basic.lean @@ -22,33 +22,15 @@ respectively. To construct an object in the category of `R`-modules from a type `M` with an instance of the `Module` typeclass, write `of R M`. There is a coercion in the other direction. +The roundtrip `↑(of R M)` is definitionally equal to `M` itself (when `M` is a type with `Module` +instance), and so is `of R ↑M` (when `M : ModuleCat R M`). -Similarly, there is a coercion from morphisms in `Module R` to linear maps. - -Porting note: the next two paragraphs should be revised. - -Unfortunately, Lean is not smart enough to see that, given an object `M : ModuleCat R`, -the expression `of R M`, where we coerce `M` to the carrier type, -is definitionally equal to `M` itself. -This means that to go the other direction, i.e., from linear maps/equivalences to (iso)morphisms -in the category of `R`-modules, we have to take care not to inadvertently end up with an -`of R M` where `M` is already an object. Hence, given `f : M →ₗ[R] N`, -* if `M N : ModuleCat R`, simply use `f`; -* if `M : ModuleCat R` and `N` is an unbundled `R`-module, use `↿f` or `asHomLeft f`; -* if `M` is an unbundled `R`-module and `N : ModuleCat R`, use `↾f` or `asHomRight f`; -* if `M` and `N` are unbundled `R`-modules, use `↟f` or `asHom f`. - -Similarly, given `f : M ≃ₗ[R] N`, use `toModuleIso`, `toModuleIso'Left`, `toModuleIso'Right` -or `toModuleIso'`, respectively. - -The arrow notations are localized, so you may have to `open ModuleCat` (or `open scoped ModuleCat`) -to use them. Note that the notation for `asHomLeft` clashes with the notation used to promote -functions between types to morphisms in the category `Type`, so to avoid confusion, it is probably a -good idea to avoid having the locales `ModuleCat` and `CategoryTheory.Type` open at the same time. - -If you get an error when trying to apply a theorem and the `convert` tactic produces goals of the -form `M = of R M`, then you probably used an incorrect variant of `asHom` or `toModuleIso`. +The morphisms are given their own type, not identified with `LinearMap`. +There is a cast from morphisms in `Module R` to linear maps, written `f.hom` (`ModuleCat.Hom.hom`). +To go from linear maps to morphisms in `Module R`, use `ModuleCat.ofHom`. +Similarly, given an isomorphism `f : M ≅ N` use `f.toLinearEquiv` and given a linear equiv +`f : M ≃ₗ[R] N`, use `f.toModuleIso`. -/ @@ -69,6 +51,7 @@ impose here that the `ℤ`-multiplication field from the module structure is def from the `isAddCommGroup` structure (contrary to what we do for all module structures in mathlib), which creates some difficulties down the road. -/ structure ModuleCat where + private mk :: /-- the underlying type of an object in `ModuleCat R` -/ carrier : Type v [isAddCommGroup : AddCommGroup carrier] @@ -88,27 +71,137 @@ instance : CoeSort (ModuleCat.{v} R) (Type v) := attribute [coe] ModuleCat.carrier +/-- The object in the category of R-algebras associated to a type equipped with the appropriate +typeclasses. This is the preferred way to construct a term of `ModuleCat R`. -/ +abbrev of (X : Type v) [AddCommGroup X] [Module R X] : ModuleCat.{v} R := + ⟨X⟩ + +lemma coe_of (X : Type v) [Ring X] [Module R X] : (of R X : Type v) = X := + rfl + +-- Ensure the roundtrips are reducibly defeq (so tactics like `rw` can see through them). +example (X : Type v) [Ring X] [Module R X] : (of R X : Type v) = X := by with_reducible rfl +example (M : ModuleCat.{v} R) : of R M = M := by with_reducible rfl + +variable {R} in +/-- The type of morphisms in `ModuleCat R`. -/ +@[ext] +structure Hom (M N : ModuleCat.{v} R) where + private mk :: + /-- The underlying linear map. -/ + hom : M →ₗ[R] N + instance moduleCategory : Category.{v, max (v+1) u} (ModuleCat.{v} R) where - Hom M N := M →ₗ[R] N - id _ := LinearMap.id - comp f g := g.comp f - id_comp _ := LinearMap.id_comp _ - comp_id _ := LinearMap.comp_id _ - assoc f g h := LinearMap.comp_assoc (f := f) (g := g) (h := h) + Hom M N := Hom M N + id _ := ⟨LinearMap.id⟩ + comp f g := ⟨g.hom.comp f.hom⟩ + +instance {M N : ModuleCat.{v} R} : CoeFun (M ⟶ N) (fun _ ↦ M → N) where + coe f := f.hom + +section + +variable {R} + +@[simp] +lemma hom_id {M : ModuleCat.{v} R} : (𝟙 M : M ⟶ M).hom = LinearMap.id := rfl + +/- Provided for rewriting. -/ +lemma id_apply (M : ModuleCat.{v} R) (x : M) : + (𝟙 M : M ⟶ M) x = x := by simp + +@[simp] +lemma hom_comp {M N O : ModuleCat.{v} R} (f : M ⟶ N) (g : N ⟶ O) : + (f ≫ g).hom = g.hom.comp f.hom := rfl + +/- Provided for rewriting. -/ +lemma comp_apply {M N O : ModuleCat.{v} R} (f : M ⟶ N) (g : N ⟶ O) (x : M) : + (f ≫ g) x = g (f x) := by simp + +@[ext] +lemma hom_ext {M N : ModuleCat.{v} R} {f g : M ⟶ N} (hf : f.hom = g.hom) : f = g := + Hom.ext hf + +lemma hom_bijective {M N : ModuleCat.{v} R} : + Function.Bijective (Hom.hom : (M ⟶ N) → (M →ₗ[R] N)) where + left f g h := by cases f; cases g; simpa using h + right f := ⟨⟨f⟩, rfl⟩ + +/-- Convenience shortcut for `ModuleCat.hom_bijective.injective`. -/ +lemma hom_injective {M N : ModuleCat.{v} R} : + Function.Injective (Hom.hom : (M ⟶ N) → (M →ₗ[R] N)) := + hom_bijective.injective + +/-- Convenience shortcut for `ModuleCat.hom_bijective.surjective`. -/ +lemma hom_surjective {M N : ModuleCat.{v} R} : + Function.Surjective (Hom.hom : (M ⟶ N) → (M →ₗ[R] N)) := + hom_bijective.surjective + +/-- Typecheck a `LinearMap` as a morphism in `ModuleCat R`. -/ +abbrev ofHom {X Y : Type v} [AddCommGroup X] [Module R X] [AddCommGroup Y] [Module R Y] + (f : X →ₗ[R] Y) : of R X ⟶ of R Y := + ⟨f⟩ + +@[deprecated (since := "2024-10-06")] alias asHom := ModuleCat.ofHom -instance {M N : ModuleCat.{v} R} : FunLike (M ⟶ N) M N := - LinearMap.instFunLike +/- Doesn't need to be `@[simp]` since the `simp` tactic applies this rewrite automatically: +`ofHom` and `hom` are reducibly equal to the constructor and projection respectively. -/ +lemma hom_ofHom {X Y : Type v} [AddCommGroup X] [Module R X] [AddCommGroup Y] + [Module R Y] (f : X →ₗ[R] Y) : (ofHom f).hom = f := rfl -instance {M N : ModuleCat.{v} R} : LinearMapClass (M ⟶ N) R M N := - LinearMap.semilinearMapClass +@[simp] +lemma ofHom_hom {M N : ModuleCat.{v} R} (f : M ⟶ N) : + ofHom (Hom.hom f) = f := rfl + +@[simp] +lemma ofHom_id {M : Type v} [AddCommGroup M] [Module R M] : ofHom LinearMap.id = 𝟙 (of R M) := rfl + +@[simp] +lemma ofHom_comp {M N O : Type v} [AddCommGroup M] [AddCommGroup N] [AddCommGroup O] [Module R M] + [Module R N] [Module R O] (f : M →ₗ[R] N) (g : N →ₗ[R] O) : + ofHom (g.comp f) = ofHom f ≫ ofHom g := + rfl + +/- Doesn't need to be `@[simp]` since `simp only` can solve this. -/ +lemma ofHom_apply {M N : Type v} [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] + (f : M →ₗ[R] N) (x : M) : ofHom f x = f x := rfl + +@[simp] +lemma inv_hom_apply {M N : ModuleCat.{v} R} (e : M ≅ N) (x : M) : e.inv (e.hom x) = x := by + rw [← comp_apply] + simp + +@[simp] +lemma hom_inv_apply {M N : ModuleCat.{v} R} (e : M ≅ N) (x : N) : e.hom (e.inv x) = x := by + rw [← comp_apply] + simp + +/-- `ModuleCat.Hom.hom` bundled as an `Equiv`. -/ +def homEquiv {M N : ModuleCat.{v} R} : (M ⟶ N) ≃ (M →ₗ[R] N) where + toFun := Hom.hom + invFun := ofHom + left_inv _ := rfl + right_inv _ := rfl + +end + +instance : Inhabited (ModuleCat R) := + ⟨of R R⟩ instance moduleConcreteCategory : ConcreteCategory.{v} (ModuleCat.{v} R) where forget := { obj := fun R => R - map := fun f => f.toFun } - forget_faithful := ⟨fun h => LinearMap.ext (fun x => by - dsimp at h - rw [h])⟩ + map := fun f => f.hom } + forget_faithful := ⟨fun h => by ext x; simpa using congrFun h x⟩ + +/- Not a `@[simp]` lemma since it will rewrite the (co)domain of maps and cause +definitional equality issues. -/ +lemma forget_obj {M : ModuleCat.{v} R} : (forget (ModuleCat.{v} R)).obj M = M := rfl + +/- Not a `@[simp]` lemma since the LHS is a categorical arrow and the RHS is a plain function. -/ +lemma forget_map {M N : ModuleCat.{v} R} (f : M ⟶ N) : + (forget (ModuleCat.{v} R)).map f = f := + rfl -- Porting note: -- One might hope these two instances would not be needed, @@ -119,18 +212,10 @@ instance {M : ModuleCat.{v} R} : AddCommGroup ((forget (ModuleCat R)).obj M) := instance {M : ModuleCat.{v} R} : Module R ((forget (ModuleCat R)).obj M) := (inferInstance : Module R M) -@[ext] -lemma ext {M N : ModuleCat.{v} R} {f₁ f₂ : M ⟶ N} (h : ∀ (x : M), f₁ x = f₂ x) : f₁ = f₂ := - DFunLike.ext _ _ h - instance hasForgetToAddCommGroup : HasForget₂ (ModuleCat R) AddCommGrp where forget₂ := { obj := fun M => AddCommGrp.of M - map := fun f => AddCommGrp.ofHom f.toAddMonoidHom } - -/-- The object in the category of R-modules associated to an R-module -/ -def of (X : Type v) [AddCommGroup X] [Module R X] : ModuleCat R := - ⟨X⟩ + map := fun f => AddCommGrp.ofHom f.hom.toAddMonoidHom } @[simp] theorem forget₂_obj (X : ModuleCat R) : @@ -148,22 +233,14 @@ theorem forget₂_obj_moduleCat_of (X : Type v) [AddCommGroup X] [Module R X] : @[simp] theorem forget₂_map (X Y : ModuleCat R) (f : X ⟶ Y) : - (forget₂ (ModuleCat R) AddCommGrp).map f = LinearMap.toAddMonoidHom f := + (forget₂ (ModuleCat R) AddCommGrp).map f = f.hom.toAddMonoidHom := rfl instance : Inhabited (ModuleCat R) := ⟨of R PUnit⟩ -instance ofUnique {X : Type v} [AddCommGroup X] [Module R X] [i : Unique X] : Unique (of R X) := - i - @[simp] theorem of_coe (X : ModuleCat R) : of R X = X := rfl --- Porting note: the simpNF linter complains, but we really need this?! --- @[simp, nolint simpNF] -theorem coe_of (X : Type v) [AddCommGroup X] [Module R X] : (of R X : Type v) = X := - rfl - variable {R} /-- Forgetting to the underlying type and then building the bundled object returns the original @@ -174,33 +251,18 @@ def ofSelfIso (M : ModuleCat R) : ModuleCat.of R M ≅ M where inv := 𝟙 M theorem isZero_of_subsingleton (M : ModuleCat R) [Subsingleton M] : IsZero M where - unique_to X := ⟨⟨⟨(0 : M →ₗ[R] X)⟩, fun f => by + unique_to X := ⟨⟨⟨ofHom (0 : M →ₗ[R] X)⟩, fun f => by ext x rw [Subsingleton.elim x (0 : M)] dsimp simp⟩⟩ - unique_from X := ⟨⟨⟨(0 : X →ₗ[R] M)⟩, fun f => by + unique_from X := ⟨⟨⟨ofHom (0 : X →ₗ[R] M)⟩, fun f => by ext x subsingleton⟩⟩ instance : HasZeroObject (ModuleCat.{v} R) := ⟨⟨of R PUnit, isZero_of_subsingleton _⟩⟩ -variable {M N U : ModuleCat.{v} R} - -@[simp] -theorem id_apply (m : M) : (𝟙 M : M → M) m = m := - rfl - -@[simp] -theorem coe_comp (f : M ⟶ N) (g : N ⟶ U) : (f ≫ g : M → U) = g ∘ f := - rfl - -theorem comp_def (f : M ⟶ N) (g : N ⟶ U) : f ≫ g = g.comp f := - rfl - -@[simp] lemma forget_map (f : M ⟶ N) : (forget (ModuleCat R)).map f = (f : M → N) := rfl - end ModuleCat variable {R} @@ -208,37 +270,23 @@ variable {X₁ X₂ : Type v} open ModuleCat -/-- Reinterpreting a linear map in the category of `R`-modules. -/ -def ModuleCat.ofHom [AddCommGroup X₁] [Module R X₁] [AddCommGroup X₂] [Module R X₂] : - (X₁ →ₗ[R] X₂) → (ModuleCat.of R X₁ ⟶ ModuleCat.of R X₂) := - id - -@[deprecated (since := "2024-12-03")] alias ModuleCat.asHom := ModuleCat.ofHom - /-- Reinterpreting a linear map in the category of `R`-modules -/ scoped[ModuleCat] notation "↟" f:1024 => ModuleCat.ofHom f -@[simp 1100] -theorem ModuleCat.ofHom_apply {R : Type u} [Ring R] {X Y : Type v} [AddCommGroup X] [Module R X] - [AddCommGroup Y] [Module R Y] (f : X →ₗ[R] Y) (x : X) : (↟ f) x = f x := - rfl - @[deprecated (since := "2024-10-06")] alias ModuleCat.asHom_apply := ModuleCat.ofHom_apply -/-- Reinterpreting a linear map in the category of `R`-modules. -/ -def ModuleCat.asHomRight [AddCommGroup X₁] [Module R X₁] {X₂ : ModuleCat.{v} R} : - (X₁ →ₗ[R] X₂) → (ModuleCat.of R X₁ ⟶ X₂) := - id +-- Since `of` and the coercion now roundtrip reducibly, we don't need to distinguish in which place +-- we need to add `of` when coercing from linear maps to morphisms. +@[deprecated ModuleCat.ofHom (since := "2024-11-29")] alias ModuleCat.asHomRight := ModuleCat.ofHom +@[deprecated ModuleCat.ofHom (since := "2024-11-29")] alias ModuleCat.asHomLeft := ModuleCat.ofHom -/-- Reinterpreting a linear map in the category of `R`-modules. -/ +/-- Reinterpreting a linear map in the category of `R`-modules. +This notation is deprecated: use `↟` instead. +-/ scoped[ModuleCat] notation "↾" f:1024 => ModuleCat.asHomRight f - -/-- Reinterpreting a linear map in the category of `R`-modules. -/ -def ModuleCat.asHomLeft {X₁ : ModuleCat.{v} R} [AddCommGroup X₂] [Module R X₂] : - (X₁ →ₗ[R] X₂) → (X₁ ⟶ ModuleCat.of R X₂) := - id - -/-- Reinterpreting a linear map in the category of `R`-modules. -/ +/-- Reinterpreting a linear map in the category of `R`-modules. +This notation is deprecated: use `↟` instead. +-/ scoped[ModuleCat] notation "↿" f:1024 => ModuleCat.asHomLeft f section @@ -247,21 +295,24 @@ section @[simps] def LinearEquiv.toModuleIso {g₁ : AddCommGroup X₁} {g₂ : AddCommGroup X₂} {m₁ : Module R X₁} {m₂ : Module R X₂} (e : X₁ ≃ₗ[R] X₂) : ModuleCat.of R X₁ ≅ ModuleCat.of R X₂ where - hom := (e : X₁ →ₗ[R] X₂) - inv := (e.symm : X₂ →ₗ[R] X₁) + hom := ofHom (e : X₁ →ₗ[R] X₂) + inv := ofHom (e.symm : X₂ →ₗ[R] X₁) hom_inv_id := by ext; apply e.left_inv inv_hom_id := by ext; apply e.right_inv /-- Build an isomorphism in the category `Module R` from a `LinearEquiv` between `Module`s. -/ +@[deprecated LinearEquiv.toModuleIso (since := "2024-11-29")] abbrev LinearEquiv.toModuleIso' {M N : ModuleCat.{v} R} (i : M ≃ₗ[R] N) : M ≅ N := i.toModuleIso /-- Build an isomorphism in the category `ModuleCat R` from a `LinearEquiv` between `Module`s. -/ +@[deprecated LinearEquiv.toModuleIso (since := "2024-11-29")] abbrev LinearEquiv.toModuleIso'Left {X₁ : ModuleCat.{v} R} [AddCommGroup X₂] [Module R X₂] (e : X₁ ≃ₗ[R] X₂) : X₁ ≅ ModuleCat.of R X₂ := e.toModuleIso /-- Build an isomorphism in the category `ModuleCat R` from a `LinearEquiv` between `Module`s. -/ +@[deprecated LinearEquiv.toModuleIso (since := "2024-11-29")] abbrev LinearEquiv.toModuleIso'Right [AddCommGroup X₁] [Module R X₁] {X₂ : ModuleCat.{v} R} (e : X₁ ≃ₗ[R] X₂) : ModuleCat.of R X₁ ≅ X₂ := e.toModuleIso @@ -270,7 +321,7 @@ namespace CategoryTheory.Iso /-- Build a `LinearEquiv` from an isomorphism in the category `ModuleCat R`. -/ def toLinearEquiv {X Y : ModuleCat R} (i : X ≅ Y) : X ≃ₗ[R] Y := - LinearEquiv.ofLinear i.hom i.inv i.inv_hom_id i.hom_inv_id + LinearEquiv.ofLinear i.hom.hom i.inv.hom (by aesop) (by aesop) end CategoryTheory.Iso @@ -286,45 +337,122 @@ end namespace ModuleCat -instance {M N : ModuleCat.{v} R} : AddCommGroup (M ⟶ N) := LinearMap.addCommGroup +section AddCommGroup + +variable {M N : ModuleCat.{v} R} + +instance : Add (M ⟶ N) where + add f g := ⟨f.hom + g.hom⟩ + +@[simp] lemma hom_add (f g : M ⟶ N) : (f + g).hom = f.hom + g.hom := rfl + +instance : Zero (M ⟶ N) where + zero := ⟨0⟩ + +@[simp] lemma hom_zero : (0 : M ⟶ N).hom = 0 := rfl + +instance : SMul ℕ (M ⟶ N) where + smul n f := ⟨n • f.hom⟩ + +@[simp] lemma hom_nsmul (n : ℕ) (f : M ⟶ N) : (n • f).hom = n • f.hom := rfl + +instance : Neg (M ⟶ N) where + neg f := ⟨-f.hom⟩ + +@[simp] lemma hom_neg (f : M ⟶ N) : (-f).hom = -f.hom := rfl + +instance : Sub (M ⟶ N) where + sub f g := ⟨f.hom - g.hom⟩ + +@[simp] lemma hom_sub (f g : M ⟶ N) : (f - g).hom = f.hom - g.hom := rfl + +instance : SMul ℤ (M ⟶ N) where + smul n f := ⟨n • f.hom⟩ + +@[simp] lemma hom_zsmul (n : ℕ) (f : M ⟶ N) : (n • f).hom = n • f.hom := rfl + +instance : AddCommGroup (M ⟶ N) := + Function.Injective.addCommGroup (Hom.hom) hom_injective + rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) + +@[simp] lemma hom_sum {ι : Type*} (f : ι → (M ⟶ N)) (s : Finset ι) : + (∑ i in s, f i).hom = ∑ i in s, (f i).hom := + map_sum ({ toFun := ModuleCat.Hom.hom, map_zero' := ModuleCat.hom_zero, map_add' := hom_add } : + (M ⟶ N) →+ (M →ₗ[R] N)) _ _ instance : Preadditive (ModuleCat.{v} R) where - add_comp P Q R f f' g := by - ext - dsimp - erw [map_add] - rfl instance forget₂_addCommGrp_additive : (forget₂ (ModuleCat.{v} R) AddCommGrp).Additive where +/-- `ModuleCat.Hom.hom` bundled as an additive equivalence. -/ +@[simps!] +def homAddEquiv : (M ⟶ N) ≃+ (M →ₗ[R] N) := + { homEquiv with + map_add' := fun _ _ => rfl } + +end AddCommGroup + +section SMul + +variable {M N : ModuleCat.{v} R} {S : Type*} [Monoid S] [DistribMulAction S N] [SMulCommClass R S N] + +instance : SMul S (M ⟶ N) where + smul c f := ⟨c • f.hom⟩ + +@[simp] lemma hom_smul (s : S) (f : M ⟶ N) : (s • f).hom = s • f.hom := rfl + +end SMul + +section Module + +variable {M N : ModuleCat.{v} R} {S : Type*} [Semiring S] [Module S N] [SMulCommClass R S N] + +instance Hom.instModule : Module S (M ⟶ N) := + Function.Injective.module S + { toFun := Hom.hom, map_zero' := hom_zero, map_add' := hom_add } + hom_injective + (fun _ _ => rfl) + +/-- `ModuleCat.Hom.hom` bundled as a linear equivalence. -/ +@[simps] +def homLinearEquiv : (M ⟶ N) ≃ₗ[S] (M →ₗ[R] N) := + { homAddEquiv with + map_smul' := fun _ _ => rfl } + +end Module + section variable {S : Type u} [CommRing S] +variable {M N : ModuleCat.{v} S} + instance : Linear S (ModuleCat.{v} S) where - homModule _ _ := LinearMap.module - smul_comp := by - intros - ext - dsimp - rw [LinearMap.smul_apply, LinearMap.smul_apply, map_smul] - rfl variable {X Y X' Y' : ModuleCat.{v} S} theorem Iso.homCongr_eq_arrowCongr (i : X ≅ X') (j : Y ≅ Y') (f : X ⟶ Y) : - Iso.homCongr i j f = LinearEquiv.arrowCongr i.toLinearEquiv j.toLinearEquiv f := + Iso.homCongr i j f = ⟨LinearEquiv.arrowCongr i.toLinearEquiv j.toLinearEquiv f.hom⟩ := rfl theorem Iso.conj_eq_conj (i : X ≅ X') (f : End X) : - Iso.conj i f = LinearEquiv.conj i.toLinearEquiv f := + Iso.conj i f = ⟨LinearEquiv.conj i.toLinearEquiv f.hom⟩ := rfl end variable (M N : ModuleCat.{v} R) +/-- `ModuleCat.Hom.hom` as an isomorphism of monoids. -/ +@[simps] +def endMulEquiv : End M ≃* (M →ₗ[R] M) where + toFun := ModuleCat.Hom.hom + invFun := ModuleCat.ofHom + map_mul' _ _ := rfl + left_inv _ := rfl + right_inv _ := rfl + /-- The scalar multiplication on an object of `ModuleCat R` considered as a morphism of rings from `R` to the endomorphisms of the underlying abelian group. -/ def smul : R →+* End ((forget₂ (ModuleCat R) AddCommGrp).obj M) where @@ -341,7 +469,7 @@ lemma smul_naturality {M N : ModuleCat.{v} R} (f : M ⟶ N) (r : R) : (forget₂ (ModuleCat R) AddCommGrp).map f ≫ N.smul r = M.smul r ≫ (forget₂ (ModuleCat R) AddCommGrp).map f := by ext x - exact (f.map_smul r x).symm + exact (f.hom.map_smul r x).symm variable (R) @@ -409,9 +537,9 @@ a morphism between the underlying objects in `AddCommGrp` and the compatibility with the scalar multiplication. -/ @[simps] def homMk : M ⟶ N where - toFun := φ - map_add' _ _ := φ.map_add _ _ - map_smul' r x := (congr_hom (hφ r) x).symm + hom.toFun := φ + hom.map_add' _ _ := φ.map_add _ _ + hom.map_smul' r x := (congr_hom (hφ r) x).symm lemma forget₂_map_homMk : (forget₂ (ModuleCat R) AddCommGrp).map (homMk φ hφ) = φ := rfl @@ -420,7 +548,7 @@ end instance : (forget (ModuleCat.{v} R)).ReflectsIsomorphisms where reflects f _ := - (inferInstance : IsIso ((LinearEquiv.mk f + (inferInstance : IsIso ((LinearEquiv.mk f.hom (asIso ((forget (ModuleCat R)).map f)).toEquiv.invFun (Equiv.left_inv _) (Equiv.right_inv _)).toModuleIso).hom) @@ -433,15 +561,45 @@ instance : (forget₂ (ModuleCat.{v} R) AddCommGrp.{v}).ReflectsIsomorphisms whe end ModuleCat +section Bilinear + +variable {R : Type*} [CommRing R] + +namespace ModuleCat + +/-- Turn a bilinear map into a homomorphism. -/ +@[simps] +def ofHom₂ {M N P : ModuleCat.{u} R} (f : M →ₗ[R] N →ₗ[R] P) : + M ⟶ of R (N ⟶ P) := + ofHom <| homLinearEquiv.symm.toLinearMap ∘ₗ f + +/-- Turn a homomorphism into a bilinear map. -/ +@[simps!] +def Hom.hom₂ {M N P : ModuleCat.{u} R} + -- We write `Hom` instead of `M ⟶ (of R (N ⟶ P))`, otherwise dot notation breaks + -- since it is expecting the type of `f` to be `ModuleCat.Hom`, not `Quiver.Hom`. + (f : Hom M (of R (N ⟶ P))) : + M →ₗ[R] N →ₗ[R] P := + Hom.hom (by convert (f ≫ ofHom homLinearEquiv.toLinearMap)) + +@[simp] lemma Hom.hom₂_ofHom₂ {M N P : ModuleCat.{u} R} (f : M →ₗ[R] N →ₗ[R] P) : + (ofHom₂ f).hom₂ = f := rfl + +@[simp] lemma ofHom₂_hom₂ {M N P : ModuleCat.{u} R} (f : M ⟶ of R (N ⟶ P)) : + ofHom₂ f.hom₂ = f := rfl + +end ModuleCat + +end Bilinear + /-! `@[simp]` lemmas for `LinearMap.comp` and categorical identities. -/ @[simp] theorem LinearMap.comp_id_moduleCat {R} [Ring R] {G : ModuleCat.{u} R} {H : Type u} [AddCommGroup H] [Module R H] (f : G →ₗ[R] H) : - f.comp (𝟙 G) = f := - Category.id_comp (ModuleCat.ofHom f) + f.comp (𝟙 G : G ⟶ G).hom = f := by simp + @[simp] theorem LinearMap.id_moduleCat_comp {R} [Ring R] {G : Type u} [AddCommGroup G] [Module R G] {H : ModuleCat.{u} R} (f : G →ₗ[R] H) : - LinearMap.comp (𝟙 H) f = f := - Category.comp_id (ModuleCat.ofHom f) + LinearMap.comp (𝟙 H : H ⟶ H).hom f = f := by simp diff --git a/Mathlib/Algebra/Category/ModuleCat/Biproducts.lean b/Mathlib/Algebra/Category/ModuleCat/Biproducts.lean index 3c026878abc0c..4387064834c7e 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Biproducts.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Biproducts.lean @@ -41,10 +41,12 @@ def binaryProductLimitCone (M N : ModuleCat.{v} R) : Limits.LimitCone (pair M N) π := { app := fun j => Discrete.casesOn j fun j => - WalkingPair.casesOn j (LinearMap.fst R M N) (LinearMap.snd R M N) + WalkingPair.casesOn j (ofHom <| LinearMap.fst R M N) (ofHom <| LinearMap.snd R M N) naturality := by rintro ⟨⟨⟩⟩ ⟨⟨⟩⟩ ⟨⟨⟨⟩⟩⟩ <;> rfl } } isLimit := - { lift := fun s => LinearMap.prod (s.π.app ⟨WalkingPair.left⟩) (s.π.app ⟨WalkingPair.right⟩) + { lift := fun s => ofHom <| LinearMap.prod + (s.π.app ⟨WalkingPair.left⟩).hom + (s.π.app ⟨WalkingPair.right⟩).hom fac := by rintro s (⟨⟩ | ⟨⟩) <;> rfl uniq := fun s m w => by simp_rw [← w ⟨WalkingPair.left⟩, ← w ⟨WalkingPair.right⟩] @@ -52,12 +54,12 @@ def binaryProductLimitCone (M N : ModuleCat.{v} R) : Limits.LimitCone (pair M N) @[simp] theorem binaryProductLimitCone_cone_π_app_left (M N : ModuleCat.{v} R) : - (binaryProductLimitCone M N).cone.π.app ⟨WalkingPair.left⟩ = LinearMap.fst R M N := + (binaryProductLimitCone M N).cone.π.app ⟨WalkingPair.left⟩ = ofHom (LinearMap.fst R M N) := rfl @[simp] theorem binaryProductLimitCone_cone_π_app_right (M N : ModuleCat.{v} R) : - (binaryProductLimitCone M N).cone.π.app ⟨WalkingPair.right⟩ = LinearMap.snd R M N := + (binaryProductLimitCone M N).cone.π.app ⟨WalkingPair.right⟩ = ofHom (LinearMap.snd R M N) := rfl /-- We verify that the biproduct in `ModuleCat R` is isomorphic to @@ -69,12 +71,12 @@ noncomputable def biprodIsoProd (M N : ModuleCat.{v} R) : @[simp, elementwise] theorem biprodIsoProd_inv_comp_fst (M N : ModuleCat.{v} R) : - (biprodIsoProd M N).inv ≫ biprod.fst = LinearMap.fst R M N := + (biprodIsoProd M N).inv ≫ biprod.fst = ofHom (LinearMap.fst R M N) := IsLimit.conePointUniqueUpToIso_inv_comp _ _ (Discrete.mk WalkingPair.left) @[simp, elementwise] theorem biprodIsoProd_inv_comp_snd (M N : ModuleCat.{v} R) : - (biprodIsoProd M N).inv ≫ biprod.snd = LinearMap.snd R M N := + (biprodIsoProd M N).inv ≫ biprod.snd = ofHom (LinearMap.snd R M N) := IsLimit.conePointUniqueUpToIso_inv_comp _ _ (Discrete.mk WalkingPair.right) namespace HasLimit @@ -85,14 +87,15 @@ variable {J : Type w} (f : J → ModuleCat.{max w v} R) to the cartesian product of those groups. -/ @[simps] -def lift (s : Fan f) : s.pt ⟶ ModuleCat.of R (∀ j, f j) where - toFun x j := s.π.app ⟨j⟩ x - map_add' x y := by - simp only [Functor.const_obj_obj, map_add] - rfl - map_smul' r x := by - simp only [Functor.const_obj_obj, map_smul] - rfl +def lift (s : Fan f) : s.pt ⟶ ModuleCat.of R (∀ j, f j) := + ofHom + { toFun := fun x j => s.π.app ⟨j⟩ x + map_add' := fun x y => by + simp only [Functor.const_obj_obj, map_add] + rfl + map_smul' := fun r x => by + simp only [Functor.const_obj_obj, map_smul] + rfl } /-- Construct limit data for a product in `ModuleCat R`, using `ModuleCat.of R (∀ j, F.obj j)`. -/ @@ -100,13 +103,12 @@ def lift (s : Fan f) : s.pt ⟶ ModuleCat.of R (∀ j, f j) where def productLimitCone : Limits.LimitCone (Discrete.functor f) where cone := { pt := ModuleCat.of R (∀ j, f j) - π := Discrete.natTrans fun j => (LinearMap.proj j.as : (∀ j, f j) →ₗ[R] f j.as) } + π := Discrete.natTrans fun j => ofHom (LinearMap.proj j.as : (∀ j, f j) →ₗ[R] f j.as) } isLimit := { lift := lift.{_, v} f fac := fun _ _ => rfl uniq := fun s m w => by - ext x - funext j + ext x j exact congr_arg (fun g : s.pt ⟶ f j => (g : s.pt → f j) x) (w ⟨j⟩) } end HasLimit @@ -124,7 +126,7 @@ noncomputable def biproductIsoPi [Finite J] (f : J → ModuleCat.{v} R) : @[simp, elementwise] theorem biproductIsoPi_inv_comp_π [Finite J] (f : J → ModuleCat.{v} R) (j : J) : - (biproductIsoPi f).inv ≫ biproduct.π f j = (LinearMap.proj j : (∀ j, f j) →ₗ[R] f j) := + (biproductIsoPi f).inv ≫ biproduct.π f j = ofHom (LinearMap.proj j : (∀ j, f j) →ₗ[R] f j) := IsLimit.conePointUniqueUpToIso_inv_comp _ _ (Discrete.mk j) end ModuleCat @@ -145,7 +147,7 @@ noncomputable def lequivProdOfRightSplitExact {f : B →ₗ[R] M} (hj : Function (exac : LinearMap.range j = LinearMap.ker g) (h : g.comp f = LinearMap.id) : (A × B) ≃ₗ[R] M := ((ShortComplex.Splitting.ofExactOfSection _ (ShortComplex.Exact.moduleCat_of_range_eq_ker (ModuleCat.ofHom j) - (ModuleCat.ofHom g) exac) (ofHom f) h + (ModuleCat.ofHom g) exac) (ofHom f) (hom_ext h) (by simpa only [ModuleCat.mono_iff_injective])).isoBinaryBiproduct ≪≫ biprodIsoProd _ _ ).symm.toLinearEquiv @@ -155,7 +157,7 @@ noncomputable def lequivProdOfLeftSplitExact {f : M →ₗ[R] A} (hg : Function. (exac : LinearMap.range j = LinearMap.ker g) (h : f.comp j = LinearMap.id) : (A × B) ≃ₗ[R] M := ((ShortComplex.Splitting.ofExactOfRetraction _ (ShortComplex.Exact.moduleCat_of_range_eq_ker (ModuleCat.ofHom j) - (ModuleCat.ofHom g) exac) (ModuleCat.ofHom f) h + (ModuleCat.ofHom g) exac) (ModuleCat.ofHom f) (hom_ext h) (by simpa only [ModuleCat.epi_iff_surjective] using hg)).isoBinaryBiproduct ≪≫ biprodIsoProd _ _).symm.toLinearEquiv diff --git a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean index 4bbcee39a959e..1d40a7cc808f9 100644 --- a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean +++ b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean @@ -56,15 +56,18 @@ variable (M : ModuleCat.{v} S) /-- Any `S`-module M is also an `R`-module via a ring homomorphism `f : R ⟶ S` by defining `r • m := f r • m` (`Module.compHom`). This is called restriction of scalars. -/ -def obj' : ModuleCat R where - carrier := M - isModule := Module.compHom M f +def obj' : ModuleCat R := + let _ := Module.compHom M f + of R M /-- Given an `S`-linear map `g : M → M'` between `S`-modules, `g` is also `R`-linear between `M` and `M'` by means of restriction of scalars. -/ def map' {M M' : ModuleCat.{v} S} (g : M ⟶ M') : obj' f M ⟶ obj' f M' := - { g with map_smul' := fun r => g.map_smul (f r) } + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(X := ...)` and `(Y := ...)`. + -- This suggests `RestrictScalars.obj'` needs to be redesigned. + ofHom (X := obj' f M) (Y := obj' f M') + { g.hom with map_smul' := fun r => g.hom.map_smul (f r) } end RestrictScalars @@ -76,24 +79,21 @@ def restrictScalars {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] (f : R → ModuleCat.{v} S ⥤ ModuleCat.{v} R where obj := RestrictScalars.obj' f map := RestrictScalars.map' f - map_id _ := LinearMap.ext fun _ => rfl - map_comp _ _ := LinearMap.ext fun _ => rfl instance {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] (f : R →+* S) : (restrictScalars.{v} f).Faithful where - map_injective h := - LinearMap.ext fun x => by simpa only using DFunLike.congr_fun h x + map_injective h := by + ext x + simpa only using DFunLike.congr_fun (ModuleCat.hom_ext_iff.mp h) x instance {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] (f : R →+* S) : (restrictScalars.{v} f).PreservesMonomorphisms where preserves _ h := by rwa [mono_iff_injective] at h ⊢ -- Porting note: this should be automatic -instance {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] {f : R →+* S} - {M : ModuleCat.{v} S} : Module R <| (restrictScalars f).obj M := - inferInstanceAs <| Module R <| RestrictScalars.obj' f M - --- Porting note: this should be automatic +-- TODO: this instance gives diamonds if `f : S →+* S`, see `PresheafOfModules.pushforward₀`. +-- The correct solution is probably to define explicit maps between `M` and +-- `(restrictScalars f).obj M`. instance {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] {f : R →+* S} {M : ModuleCat.{v} S} : Module S <| (restrictScalars f).obj M := inferInstanceAs <| Module S M @@ -105,20 +105,18 @@ theorem restrictScalars.map_apply {R : Type u₁} {S : Type u₂} [Ring R] [Ring @[simp] theorem restrictScalars.smul_def {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] (f : R →+* S) - {M : ModuleCat.{v} S} (r : R) (m : (restrictScalars f).obj M) : r • m = (f r • m : M) := + {M : ModuleCat.{v} S} (r : R) (m : (restrictScalars f).obj M) : r • m = f r • show M from m := rfl theorem restrictScalars.smul_def' {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] (f : R →+* S) {M : ModuleCat.{v} S} (r : R) (m : M) : - -- Porting note: clumsy way to coerce - let m' : (restrictScalars f).obj M := m - (r • m' : (restrictScalars f).obj M) = (f r • m : M) := + r • (show (restrictScalars f).obj M from m) = f r • m := rfl instance (priority := 100) sMulCommClass_mk {R : Type u₁} {S : Type u₂} [Ring R] [CommRing S] (f : R →+* S) (M : Type v) [I : AddCommGroup M] [Module S M] : - haveI : SMul R M := (RestrictScalars.obj' f (ModuleCat.mk M)).isModule.toSMul + haveI : SMul R M := (RestrictScalars.obj' f (ModuleCat.of S M)).isModule.toSMul SMulCommClass R S M := @SMulCommClass.mk R S M (_) _ fun r s m => (by simp [← mul_smul, mul_comm] : f r • s • m = s • f r • m) @@ -129,14 +127,16 @@ morphisms `M ⟶ (ModuleCat.restrictScalars f).obj N`. -/ def semilinearMapAddEquiv {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] (f : R →+* S) (M : ModuleCat.{v} R) (N : ModuleCat.{v} S) : (M →ₛₗ[f] N) ≃+ (M ⟶ (ModuleCat.restrictScalars f).obj N) where - toFun g := + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(Y := ...)`. + -- This suggests `restrictScalars` needs to be redesigned. + toFun g := ofHom (Y := (ModuleCat.restrictScalars f).obj N) <| { toFun := g map_add' := by simp map_smul' := by simp } invFun g := { toFun := g map_add' := by simp - map_smul' := g.map_smul } + map_smul' := g.hom.map_smul } left_inv _ := rfl right_inv _ := rfl map_add' _ _ := rfl @@ -149,7 +149,7 @@ variable {R : Type u₁} [Ring R] (f : R →+* R) to `M`. -/ def restrictScalarsId'App (hf : f = RingHom.id R) (M : ModuleCat R) : (restrictScalars f).obj M ≅ M := - LinearEquiv.toModuleIso' <| + LinearEquiv.toModuleIso <| @AddEquiv.toLinearEquiv _ _ _ _ _ _ (((restrictScalars f).obj M).isModule) _ (by rfl) (fun r x ↦ by subst hf; rfl) @@ -198,7 +198,11 @@ variable {R₁ : Type u₁} {R₂ : Type u₂} {R₃ : Type u₃} [Ring R₁] [R identifies to successively restricting scalars. -/ def restrictScalarsComp'App (hgf : gf = g.comp f) (M : ModuleCat R₃) : (restrictScalars gf).obj M ≅ (restrictScalars f).obj ((restrictScalars g).obj M) := - (AddEquiv.toLinearEquiv (by rfl) (fun r x ↦ by subst hgf; rfl)).toModuleIso' + (AddEquiv.toLinearEquiv + (M := ↑((restrictScalars gf).obj M)) + (M₂ := ↑((restrictScalars f).obj ((restrictScalars g).obj M))) + (by rfl) + (fun r x ↦ by subst hgf; rfl)).toModuleIso variable (hgf : gf = g.comp f) @@ -243,10 +247,14 @@ def restrictScalarsEquivalenceOfRingEquiv {R S} [Ring R] [Ring S] (e : R ≃+* S ModuleCat S ≌ ModuleCat R where functor := ModuleCat.restrictScalars e.toRingHom inverse := ModuleCat.restrictScalars e.symm - unitIso := NatIso.ofComponents (fun M ↦ LinearEquiv.toModuleIso' + unitIso := NatIso.ofComponents (fun M ↦ LinearEquiv.toModuleIso + (X₁ := M) + (X₂ := (restrictScalars e.symm.toRingHom).obj ((restrictScalars e.toRingHom).obj M)) { __ := AddEquiv.refl M map_smul' := fun s m ↦ congr_arg (· • m) (e.right_inv s).symm }) (by intros; rfl) - counitIso := NatIso.ofComponents (fun M ↦ LinearEquiv.toModuleIso' + counitIso := NatIso.ofComponents (fun M ↦ LinearEquiv.toModuleIso + (X₁ := (restrictScalars e.toRingHom).obj ((restrictScalars e.symm.toRingHom).obj M)) + (X₂ := M) { __ := AddEquiv.refl M map_smul' := fun r _ ↦ congr_arg (· • (_ : M)) (e.left_inv r)}) (by intros; rfl) functor_unitIso_comp := by intros; rfl @@ -279,33 +287,26 @@ variable (M : ModuleCat.{v} R) /-- Extension of scalars turn an `R`-module into `S`-module by M ↦ S ⨂ M -/ def obj' : ModuleCat S := - ⟨TensorProduct R ((restrictScalars f).obj ⟨S⟩) M⟩ + of _ (TensorProduct R ((restrictScalars f).obj (of _ S)) M) /-- Extension of scalars is a functor where an `R`-module `M` is sent to `S ⊗ M` and `l : M1 ⟶ M2` is sent to `s ⊗ m ↦ s ⊗ l m` -/ def map' {M1 M2 : ModuleCat.{v} R} (l : M1 ⟶ M2) : obj' f M1 ⟶ obj' f M2 := - by-- The "by apply" part makes this require 75% fewer heartbeats to process (https://github.com/leanprover-community/mathlib4/pull/16371). - apply @LinearMap.baseChange R S M1 M2 _ _ ((algebraMap S _).comp f).toAlgebra _ _ _ _ l - -theorem map'_id {M : ModuleCat.{v} R} : map' f (𝟙 M) = 𝟙 _ := - LinearMap.ext fun x : obj' f M => by - dsimp only [map'] - rw [ModuleCat.id_apply] -- Porting note: this got put in the dsimp by mathport - induction x using TensorProduct.induction_on with - | zero => rw [map_zero] - | tmul => -- Porting note: issues with synthesizing Algebra R S - erw [@LinearMap.baseChange_tmul R S M M _ _ (_), ModuleCat.id_apply] - | add _ _ ihx ihy => rw [map_add, ihx, ihy] + ofHom (@LinearMap.baseChange R S M1 M2 _ _ ((algebraMap S _).comp f).toAlgebra _ _ _ _ l.hom) + +theorem map'_id {M : ModuleCat.{v} R} : map' f (𝟙 M) = 𝟙 _ := by + ext x + simp [map'] theorem map'_comp {M₁ M₂ M₃ : ModuleCat.{v} R} (l₁₂ : M₁ ⟶ M₂) (l₂₃ : M₂ ⟶ M₃) : - map' f (l₁₂ ≫ l₂₃) = map' f l₁₂ ≫ map' f l₂₃ := - LinearMap.ext fun x : obj' f M₁ => by - dsimp only [map'] - induction x using TensorProduct.induction_on with - | zero => rfl - | tmul => rfl - | add _ _ ihx ihy => rw [map_add, map_add, ihx, ihy] + map' f (l₁₂ ≫ l₂₃) = map' f l₁₂ ≫ map' f l₂₃ := by + ext x + dsimp only [map'] + induction x using TensorProduct.induction_on with + | zero => rfl + | tmul => rfl + | add _ _ ihx ihy => rw [map_add, map_add, ihx, ihy] end ExtendScalars @@ -352,7 +353,7 @@ variable (M : Type v) [AddCommMonoid M] [Module R M] /-- Given an `R`-module M, consider Hom(S, M) -- the `R`-linear maps between S (as an `R`-module by means of restriction of scalars) and M. `S` acts on Hom(S, M) by `s • g = x ↦ g (x • s)` -/ -instance hasSMul : SMul S <| (restrictScalars f).obj ⟨S⟩ →ₗ[R] M where +instance hasSMul : SMul S <| (restrictScalars f).obj (of _ S) →ₗ[R] M where smul s g := { toFun := fun s' : S => g (s' * s : S) map_add' := fun x y : S => by dsimp; rw [add_mul, map_add] @@ -363,16 +364,16 @@ instance hasSMul : SMul S <| (restrictScalars f).obj ⟨S⟩ →ₗ[R] M where erw [smul_eq_mul, smul_eq_mul, mul_assoc] } @[simp] -theorem smul_apply' (s : S) (g : (restrictScalars f).obj ⟨S⟩ →ₗ[R] M) (s' : S) : +theorem smul_apply' (s : S) (g : (restrictScalars f).obj (of _ S) →ₗ[R] M) (s' : S) : (s • g) s' = g (s' * s : S) := rfl -instance mulAction : MulAction S <| (restrictScalars f).obj ⟨S⟩ →ₗ[R] M := +instance mulAction : MulAction S <| (restrictScalars f).obj (of _ S) →ₗ[R] M := { CoextendScalars.hasSMul f _ with one_smul := fun g => LinearMap.ext fun s : S => by simp mul_smul := fun (s t : S) g => LinearMap.ext fun x : S => by simp [mul_assoc] } -instance distribMulAction : DistribMulAction S <| (restrictScalars f).obj ⟨S⟩ →ₗ[R] M := +instance distribMulAction : DistribMulAction S <| (restrictScalars f).obj (of _ S) →ₗ[R] M := { CoextendScalars.mulAction f _ with smul_add := fun s g h => LinearMap.ext fun _ : S => by simp smul_zero := fun _ => LinearMap.ext fun _ : S => by simp } @@ -380,7 +381,7 @@ instance distribMulAction : DistribMulAction S <| (restrictScalars f).obj ⟨S /-- `S` acts on Hom(S, M) by `s • g = x ↦ g (x • s)`, this action defines an `S`-module structure on Hom(S, M). -/ -instance isModule : Module S <| (restrictScalars f).obj ⟨S⟩ →ₗ[R] M := +instance isModule : Module S <| (restrictScalars f).obj (of _ S) →ₗ[R] M := { CoextendScalars.distribMulAction f _ with add_smul := fun s1 s2 g => LinearMap.ext fun x : S => by simp [mul_add, LinearMap.map_add] zero_smul := fun g => LinearMap.ext fun x : S => by simp [LinearMap.map_zero] } @@ -392,18 +393,19 @@ variable (M : ModuleCat.{v} R) /-- If `M` is an `R`-module, then the set of `R`-linear maps `S →ₗ[R] M` is an `S`-module with scalar multiplication defined by `s • l := x ↦ l (x • s)`-/ def obj' : ModuleCat S := - ⟨(restrictScalars f).obj ⟨S⟩ →ₗ[R] M⟩ + of _ ((restrictScalars f).obj (of _ S) →ₗ[R] M) -instance : CoeFun (obj' f M) fun _ => S → M where coe g := g.toFun +instance : CoeFun (obj' f M) fun _ => S → M := + inferInstanceAs <| CoeFun ((restrictScalars f).obj (of _ S) →ₗ[R] M) _ /-- If `M, M'` are `R`-modules, then any `R`-linear map `g : M ⟶ M'` induces an `S`-linear map `(S →ₗ[R] M) ⟶ (S →ₗ[R] M')` defined by `h ↦ g ∘ h`-/ @[simps] -def map' {M M' : ModuleCat R} (g : M ⟶ M') : obj' f M ⟶ obj' f M' where - toFun h := g.comp h - map_add' _ _ := LinearMap.comp_add _ _ _ - map_smul' s h := LinearMap.ext fun t : S => by dsimp; rw [smul_apply',smul_apply']; simp - -- Porting note: smul_apply' not working in simp +def map' {M M' : ModuleCat R} (g : M ⟶ M') : obj' f M ⟶ obj' f M' := + ofHom + { toFun := fun h => g.hom.comp h + map_add' := fun _ _ => LinearMap.comp_add _ _ _ + map_smul' := fun s h => by ext; simp } end CoextendScalars @@ -416,14 +418,13 @@ def coextendScalars {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] (f : R → ModuleCat R ⥤ ModuleCat S where obj := CoextendScalars.obj' f map := CoextendScalars.map' f - map_id _ := LinearMap.ext fun _ => LinearMap.ext fun _ => rfl - map_comp _ _ := LinearMap.ext fun _ => LinearMap.ext fun _ => rfl + map_id _ := by ext; rfl + map_comp _ _ := by ext; rfl namespace CoextendScalars variable {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] (f : R →+* S) --- Porting note: this coercion doesn't line up well with task below instance (M : ModuleCat R) : CoeFun ((coextendScalars f).obj M) fun _ => S → M := inferInstanceAs <| CoeFun (CoextendScalars.obj' f M) _ @@ -445,51 +446,50 @@ variable {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] (f : R →+* S) /-- Given `R`-module X and `S`-module Y, any `g : (restrictScalars f).obj Y ⟶ X` corresponds to `Y ⟶ (coextendScalars f).obj X` by sending `y ↦ (s ↦ g (s • y))` -/ -@[simps apply_apply] def HomEquiv.fromRestriction {X : ModuleCat R} {Y : ModuleCat S} - (g : (restrictScalars f).obj Y ⟶ X) : Y ⟶ (coextendScalars f).obj X where - toFun := fun y : Y => - { toFun := fun s : S => g <| (s • y : Y) - map_add' := fun s1 s2 : S => by simp only [add_smul]; rw [LinearMap.map_add] - map_smul' := fun r (s : S) => by - -- Porting note: dsimp clears out some rw's but less eager to apply others with Lean 4 - dsimp - rw [← g.map_smul] - erw [smul_eq_mul, mul_smul] - rfl} - map_add' := fun y1 y2 : Y => - LinearMap.ext fun s : S => by - -- Porting note: double dsimp seems odd - dsimp only [id_eq, eq_mpr_eq_cast, AddHom.toFun_eq_coe, AddHom.coe_mk, RingHom.id_apply, - RingHom.toMonoidHom_eq_coe, OneHom.toFun_eq_coe, - MonoidHom.toOneHom_coe, MonoidHom.coe_coe, ZeroHom.coe_mk, smul_eq_mul, cast_eq, - LinearMap.coe_mk] - rw [LinearMap.add_apply, LinearMap.coe_mk, LinearMap.coe_mk] - dsimp only [AddHom.coe_mk] - rw [smul_add, map_add] - map_smul' := fun (s : S) (y : Y) => LinearMap.ext fun t : S => by - -- Porting note: used to be simp [mul_smul] - simp only [LinearMap.coe_mk, AddHom.coe_mk, RingHom.id_apply] - rw [ModuleCat.CoextendScalars.smul_apply', LinearMap.coe_mk] - dsimp - rw [mul_smul] + (g : (restrictScalars f).obj Y ⟶ X) : Y ⟶ (coextendScalars f).obj X := + ofHom + { toFun := fun y : Y => + { toFun := fun s : S => g <| (s • y : Y) + map_add' := fun s1 s2 : S => by simp only [add_smul]; rw [LinearMap.map_add] + map_smul' := fun r (s : S) => by + -- Porting note: dsimp clears out some rw's but less eager to apply others with Lean 4 + dsimp + rw [← g.hom.map_smul] + erw [smul_eq_mul, mul_smul] + rfl } + map_add' := fun y1 y2 : Y => + LinearMap.ext fun s : S => by + simp [smul_add, map_add] + map_smul' := fun (s : S) (y : Y) => LinearMap.ext fun t : S => by + simp [mul_smul] } + +/-- This should be autogenerated by `@[simps]` but we need to give `s` the correct type here. -/ +@[simp] lemma HomEquiv.fromRestriction_hom_apply_apply {X : ModuleCat R} {Y : ModuleCat S} + (g : (restrictScalars f).obj Y ⟶ X) (y) (s : S) : + (HomEquiv.fromRestriction f g).hom y s = g (s • y) := rfl /-- Given `R`-module X and `S`-module Y, any `g : Y ⟶ (coextendScalars f).obj X` corresponds to `(restrictScalars f).obj Y ⟶ X` by `y ↦ g y 1` -/ -@[simps apply] def HomEquiv.toRestriction {X Y} (g : Y ⟶ (coextendScalars f).obj X) : - (restrictScalars f).obj Y ⟶ X where - toFun := fun y : Y => (g y) (1 : S) - map_add' x y := by dsimp; rw [g.map_add, LinearMap.add_apply] - map_smul' r (y : Y) := by - dsimp - rw [← LinearMap.map_smul] - erw [smul_eq_mul, mul_one, LinearMap.map_smul] - -- Porting note: should probably change CoeFun for obj above - rw [← LinearMap.coe_toAddHom, ← AddHom.toFun_eq_coe] - rw [CoextendScalars.smul_apply (s := f r) (g := g y) (s' := 1), one_mul] - simp + (restrictScalars f).obj Y ⟶ X := + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(X := ...)`. + -- This suggests `restrictScalars` needs to be redesigned. + ofHom (X := (restrictScalars f).obj Y) + { toFun := fun y : Y => (g y) (1 : S) + map_add' := fun x y => by dsimp; rw [g.hom.map_add, LinearMap.add_apply] + map_smul' := fun r (y : Y) => by + dsimp + rw [← LinearMap.map_smul] + erw [smul_eq_mul, mul_one, LinearMap.map_smul] + rw [CoextendScalars.smul_apply (s := f r) (g := g y) (s' := 1), one_mul] + simp } + +/-- This should be autogenerated by `@[simps]` but we need to give `1` the correct type here. -/ +@[simp] lemma HomEquiv.toRestriction_hom_apply {X : ModuleCat R} {Y : ModuleCat S} + (g : Y ⟶ (coextendScalars f).obj X) (y) : + (HomEquiv.toRestriction f g).hom y = g.hom y (1 : S) := rfl -- Porting note: add to address timeout in unit' /-- Auxiliary definition for `unit'` -/ @@ -521,17 +521,14 @@ def app' (Y : ModuleCat S) : Y →ₗ[S] (restrictScalars f ⋙ coextendScalars The natural transformation from identity functor to the composition of restriction and coextension of scalars. -/ --- @[simps] Porting note: not in normal form and not used +@[simps] protected def unit' : 𝟭 (ModuleCat S) ⟶ restrictScalars f ⋙ coextendScalars f where - app Y := app' f Y + app Y := ofHom (app' f Y) naturality Y Y' g := - LinearMap.ext fun y : Y => LinearMap.ext fun s : S => by + hom_ext <| LinearMap.ext fun y : Y => LinearMap.ext fun s : S => by -- Porting note (https://github.com/leanprover-community/mathlib4/pull/10745): previously simp [CoextendScalars.map_apply] - simp only [ModuleCat.coe_comp, Functor.id_map, Functor.id_obj, Functor.comp_obj, - Functor.comp_map] - rw [coe_comp, coe_comp, Function.comp, Function.comp] - conv_rhs => rw [← LinearMap.coe_toAddHom, ← AddHom.toFun_eq_coe] - rw [CoextendScalars.map_apply, AddHom.toFun_eq_coe, LinearMap.coe_toAddHom, + simp only [ModuleCat.hom_comp, Functor.id_map, Functor.id_obj, Functor.comp_obj, + Functor.comp_map, LinearMap.coe_comp, Function.comp, CoextendScalars.map_apply, restrictScalars.map_apply f] change s • (g y) = g (s • y) rw [map_smul] @@ -539,24 +536,21 @@ protected def unit' : 𝟭 (ModuleCat S) ⟶ restrictScalars f ⋙ coextendScala /-- The natural transformation from the composition of coextension and restriction of scalars to identity functor. -/ --- @[simps] Porting note: not in normal form and not used +@[simps] protected def counit' : coextendScalars f ⋙ restrictScalars f ⟶ 𝟭 (ModuleCat R) where - app X := + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(X := ...)`. + -- This suggests `restrictScalars` needs to be redesigned. + app X := ofHom (X := (restrictScalars f).obj ((coextendScalars f).obj X)) { toFun := fun g => g.toFun (1 : S) map_add' := fun x1 x2 => by dsimp rw [LinearMap.add_apply] map_smul' := fun r (g : (restrictScalars f).obj ((coextendScalars f).obj X)) => by dsimp - rw [← LinearMap.coe_toAddHom, ← AddHom.toFun_eq_coe] - rw [CoextendScalars.smul_apply (s := f r) (g := g) (s' := 1), one_mul, ← LinearMap.map_smul] - rw [← LinearMap.coe_toAddHom, ← AddHom.toFun_eq_coe] + rw [CoextendScalars.smul_apply, one_mul, ← LinearMap.map_smul] congr change f r = (f r) • (1 : S) rw [smul_eq_mul (a := f r) (a' := 1), mul_one] } - naturality X X' g := LinearMap.ext fun h => by - rw [ModuleCat.coe_comp] - rfl end RestrictionCoextensionAdj @@ -569,26 +563,18 @@ def restrictCoextendScalarsAdj {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] homEquiv := fun X Y ↦ { toFun := RestrictionCoextensionAdj.HomEquiv.fromRestriction.{u₁,u₂,v} f invFun := RestrictionCoextensionAdj.HomEquiv.toRestriction.{u₁,u₂,v} f - left_inv := fun g => LinearMap.ext fun x : X => by - -- Porting note (https://github.com/leanprover-community/mathlib4/pull/10745): once just simp - rw [RestrictionCoextensionAdj.HomEquiv.toRestriction_apply, AddHom.toFun_eq_coe, - LinearMap.coe_toAddHom, RestrictionCoextensionAdj.HomEquiv.fromRestriction_apply_apply, - one_smul] - right_inv := fun g => LinearMap.ext fun x => LinearMap.ext fun s : S => by + left_inv := fun g => by ext; simp + right_inv := fun g => hom_ext <| LinearMap.ext fun x => LinearMap.ext fun s : S => by -- Porting note (https://github.com/leanprover-community/mathlib4/pull/10745): once just simp - rw [RestrictionCoextensionAdj.HomEquiv.fromRestriction_apply_apply, - RestrictionCoextensionAdj.HomEquiv.toRestriction_apply, AddHom.toFun_eq_coe, - LinearMap.coe_toAddHom, LinearMap.map_smulₛₗ, RingHom.id_apply, - CoextendScalars.smul_apply', one_mul] } + rw [RestrictionCoextensionAdj.HomEquiv.fromRestriction_hom_apply_apply, + RestrictionCoextensionAdj.HomEquiv.toRestriction_hom_apply, LinearMap.map_smulₛₗ, + RingHom.id_apply, CoextendScalars.smul_apply', one_mul] } unit := RestrictionCoextensionAdj.unit'.{u₁,u₂,v} f counit := RestrictionCoextensionAdj.counit'.{u₁,u₂,v} f - homEquiv_unit := LinearMap.ext fun _ => rfl - homEquiv_counit := fun {X Y g} => LinearMap.ext <| by - -- Porting note (https://github.com/leanprover-community/mathlib4/pull/10745): previously simp [RestrictionCoextensionAdj.counit'] - intro x; dsimp - rw [coe_comp, Function.comp] - change _ = (((restrictScalars f).map g) x).toFun (1 : S) - rw [AddHom.toFun_eq_coe, LinearMap.coe_toAddHom, restrictScalars.map_apply] } + homEquiv_unit := hom_ext <| LinearMap.ext fun _ => rfl + homEquiv_counit := fun {X Y g} => by + ext + simp [RestrictionCoextensionAdj.counit'] } instance {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] (f : R →+* S) : (restrictScalars.{max u₂ w} f).IsLeftAdjoint := @@ -611,17 +597,20 @@ Given `R`-module X and `S`-module Y and a map `g : (extendScalars f).obj X ⟶ Y map `S ⨂ X → Y`, there is a `X ⟶ (restrictScalars f).obj Y`, i.e. `R`-linear map `X ⟶ Y` by `x ↦ g (1 ⊗ x)`. -/ -@[simps apply] +@[simps hom_apply] def HomEquiv.toRestrictScalars {X Y} (g : (extendScalars f).obj X ⟶ Y) : - X ⟶ (restrictScalars f).obj Y where - toFun x := g <| (1 : S)⊗ₜ[R,f]x - map_add' _ _ := by dsimp; rw [tmul_add, map_add] - map_smul' r x := by - letI : Module R S := Module.compHom S f - letI : Module R Y := Module.compHom Y f - dsimp - erw [RestrictScalars.smul_def, ← LinearMap.map_smul, tmul_smul] - congr + X ⟶ (restrictScalars f).obj Y := + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(Y := ...)`. + -- This suggests `restrictScalars` needs to be redesigned. + ofHom (Y := (restrictScalars f).obj Y) + { toFun := fun x => g <| (1 : S)⊗ₜ[R,f]x + map_add' := fun _ _ => by dsimp; rw [tmul_add, map_add] + map_smul' := fun r s => by + letI : Module R S := Module.compHom S f + letI : Module R Y := Module.compHom Y f + dsimp + erw [RestrictScalars.smul_def, ← LinearMap.map_smul, tmul_smul] + congr } -- Porting note: forced to break apart fromExtendScalars due to timeouts /-- @@ -636,7 +625,7 @@ def HomEquiv.evalAt {X : ModuleCat R} {Y : ModuleCat S} (s : S) map_add' := by intros dsimp only - rw [map_add,smul_add] } + rw [map_add, smul_add] } (by intros r x rw [AddHom.toFun_eq_coe, AddHom.coe_mk, RingHom.id_apply, @@ -647,11 +636,11 @@ Given `R`-module X and `S`-module Y and a map `X ⟶ (restrictScalars f).obj Y`, `X ⟶ Y`, there is a map `(extend_scalars f).obj X ⟶ Y`, i.e `S`-linear map `S ⨂ X → Y` by `s ⊗ x ↦ s • g x`. -/ -@[simps apply] +@[simps hom_apply] def HomEquiv.fromExtendScalars {X Y} (g : X ⟶ (restrictScalars f).obj Y) : (extendScalars f).obj X ⟶ Y := by letI m1 : Module R S := Module.compHom S f; letI m2 : Module R Y := Module.compHom Y f - refine {toFun := fun z => TensorProduct.lift ?_ z, map_add' := ?_, map_smul' := ?_} + refine ofHom {toFun := fun z => TensorProduct.lift ?_ z, map_add' := ?_, map_smul' := ?_} · refine {toFun := fun s => HomEquiv.evalAt f s g, map_add' := fun (s₁ s₂ : S) => ?_, map_smul' := fun (r : R) (s : S) => ?_} @@ -685,6 +674,7 @@ def homEquiv {X Y} : invFun := HomEquiv.fromExtendScalars.{u₁,u₂,v} f left_inv g := by letI m1 : Module R S := Module.compHom S f; letI m2 : Module R Y := Module.compHom Y f + apply hom_ext apply LinearMap.ext; intro z induction z using TensorProduct.induction_on with | zero => rw [map_zero, map_zero] @@ -698,9 +688,11 @@ def homEquiv {X Y} : | add _ _ ih1 ih2 => rw [map_add, map_add, ih1, ih2] right_inv g := by letI m1 : Module R S := Module.compHom S f; letI m2 : Module R Y := Module.compHom Y f - apply LinearMap.ext; intro x - rw [HomEquiv.toRestrictScalars_apply, HomEquiv.fromExtendScalars_apply, lift.tmul, - LinearMap.coe_mk, LinearMap.coe_mk] + ext x + rw [HomEquiv.toRestrictScalars_hom_apply] + -- This needs to be `erw` because of some unfolding in `fromExtendScalars` + erw [HomEquiv.fromExtendScalars_hom_apply] + rw [lift.tmul, LinearMap.coe_mk, LinearMap.coe_mk] dsimp rw [one_smul] @@ -708,13 +700,16 @@ def homEquiv {X Y} : For any `R`-module X, there is a natural `R`-linear map from `X` to `X ⨂ S` by sending `x ↦ x ⊗ 1` -/ -- @[simps] Porting note: not in normal form and not used -def Unit.map {X} : X ⟶ (extendScalars f ⋙ restrictScalars f).obj X where - toFun x := (1 : S)⊗ₜ[R,f]x - map_add' x x' := by dsimp; rw [TensorProduct.tmul_add] - map_smul' r x := by - letI m1 : Module R S := Module.compHom S f - -- Porting note: used to be rfl - dsimp; rw [← TensorProduct.smul_tmul,TensorProduct.smul_tmul'] +def Unit.map {X} : X ⟶ (extendScalars f ⋙ restrictScalars f).obj X := + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(Y := ...)`. + -- This suggests `restrictScalars` needs to be redesigned. + ofHom (Y := (extendScalars f ⋙ restrictScalars f).obj X) + { toFun := fun x => (1 : S)⊗ₜ[R,f]x + map_add' := fun x x' => by dsimp; rw [TensorProduct.tmul_add] + map_smul' := fun r x => by + letI m1 : Module R S := Module.compHom S f + -- Porting note: used to be rfl + dsimp; rw [← TensorProduct.smul_tmul,TensorProduct.smul_tmul'] } /-- The natural transformation from identity functor on `R`-module to the composition of extension and @@ -726,45 +721,41 @@ def unit : 𝟭 (ModuleCat R) ⟶ extendScalars f ⋙ restrictScalars.{max v u /-- For any `S`-module Y, there is a natural `R`-linear map from `S ⨂ Y` to `Y` by `s ⊗ y ↦ s • y` -/ -@[simps apply] -def Counit.map {Y} : (restrictScalars f ⋙ extendScalars f).obj Y ⟶ Y where - toFun := - letI m1 : Module R S := Module.compHom S f - letI m2 : Module R Y := Module.compHom Y f - TensorProduct.lift - { toFun := fun s : S => - { toFun := fun y : Y => s • y, - map_add' := smul_add _ - map_smul' := fun r y => by - change s • f r • y = f r • s • y - rw [← mul_smul, mul_comm, mul_smul] }, - map_add' := fun s₁ s₂ => by - ext y - change (s₁ + s₂) • y = s₁ • y + s₂ • y - rw [add_smul] - map_smul' := fun r s => by - ext y - change (f r • s) • y = (f r) • s • y - rw [smul_eq_mul, mul_smul] } - map_add' _ _ := by rw [map_add] - map_smul' s z := by - letI m1 : Module R S := Module.compHom S f - letI m2 : Module R Y := Module.compHom Y f - dsimp only - induction z using TensorProduct.induction_on with - | zero => rw [smul_zero, map_zero, smul_zero] - | tmul s' y => - rw [ExtendScalars.smul_tmul, LinearMap.coe_mk] - erw [TensorProduct.lift.tmul, TensorProduct.lift.tmul] - set s' : S := s' - change (s * s') • y = s • s' • y - rw [mul_smul] - | add _ _ ih1 ih2 => rw [smul_add, map_add, map_add, ih1, ih2, smul_add] - --- Porting note: this file has to probably be reworked when --- coercions and instance synthesis are fixed for concrete categories --- so I say nolint now and move on -attribute [nolint simpNF] Counit.map_apply +@[simps hom_apply] +def Counit.map {Y} : (restrictScalars f ⋙ extendScalars f).obj Y ⟶ Y := + ofHom + { toFun := + letI m1 : Module R S := Module.compHom S f + letI m2 : Module R Y := Module.compHom Y f + TensorProduct.lift + { toFun := fun s : S => + { toFun := fun y : Y => s • y, + map_add' := smul_add _ + map_smul' := fun r y => by + change s • f r • y = f r • s • y + rw [← mul_smul, mul_comm, mul_smul] }, + map_add' := fun s₁ s₂ => by + ext y + change (s₁ + s₂) • y = s₁ • y + s₂ • y + rw [add_smul] + map_smul' := fun r s => by + ext y + change (f r • s) • y = (f r) • s • y + rw [smul_eq_mul, mul_smul] } + map_add' := fun _ _ => by rw [map_add] + map_smul' := fun s z => by + letI m1 : Module R S := Module.compHom S f + letI m2 : Module R Y := Module.compHom Y f + dsimp only + induction z using TensorProduct.induction_on with + | zero => rw [smul_zero, map_zero, smul_zero] + | tmul s' y => + rw [ExtendScalars.smul_tmul, LinearMap.coe_mk] + erw [TensorProduct.lift.tmul, TensorProduct.lift.tmul] + set s' : S := s' + change (s * s') • y = s • s' • y + rw [mul_smul] + | add _ _ ih1 ih2 => rw [smul_add, map_add, map_add, ih1, ih2, smul_add] } /-- The natural transformation from the composition of restriction and extension of scalars to the identity functor on `S`-module. @@ -777,15 +768,13 @@ def counit : restrictScalars.{max v u₂,u₁,u₂} f ⋙ extendScalars f ⟶ letI m1 : Module R S := Module.compHom S f letI m2 : Module R Y := Module.compHom Y f letI m2 : Module R Y' := Module.compHom Y' f - apply LinearMap.ext; intro z + ext z induction z using TensorProduct.induction_on with | zero => rw [map_zero, map_zero] | tmul s' y => dsimp - rw [ModuleCat.coe_comp, ModuleCat.coe_comp, Function.comp_apply, Function.comp_apply, - ExtendScalars.map_tmul, restrictScalars.map_apply] -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [Counit.map_apply] + erw [Counit.map_hom_apply] rw [lift.tmul, LinearMap.coe_mk, LinearMap.coe_mk] set s' : S := s' change s' • g y = g (s' • y) @@ -803,20 +792,18 @@ def extendRestrictScalarsAdj {R : Type u₁} {S : Type u₂} [CommRing R] [CommR homEquiv := fun _ _ ↦ ExtendRestrictScalarsAdj.homEquiv.{v,u₁,u₂} f unit := ExtendRestrictScalarsAdj.unit.{v,u₁,u₂} f counit := ExtendRestrictScalarsAdj.counit.{v,u₁,u₂} f - homEquiv_unit := fun {X Y g} ↦ LinearMap.ext fun x => by + homEquiv_unit := fun {X Y g} ↦ hom_ext <| LinearMap.ext fun x => by dsimp - rw [ModuleCat.coe_comp, Function.comp_apply, restrictScalars.map_apply] rfl - homEquiv_counit := fun {X Y g} ↦ LinearMap.ext fun x => by + homEquiv_counit := fun {X Y g} ↦ hom_ext <| LinearMap.ext fun x => by induction x using TensorProduct.induction_on with | zero => rw [map_zero, map_zero] | tmul => rw [ExtendRestrictScalarsAdj.homEquiv_symm_apply] dsimp - rw [ModuleCat.coe_comp, Function.comp_apply] -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [ExtendRestrictScalarsAdj.Counit.map_apply] - dsimp + erw [ExtendRestrictScalarsAdj.Counit.map_hom_apply, + ExtendRestrictScalarsAdj.HomEquiv.fromExtendScalars_hom_apply] | add => rw [map_add, map_add]; congr 1 } instance {R : Type u₁} {S : Type u₂} [CommRing R] [CommRing S] (f : R →+* S) : diff --git a/Mathlib/Algebra/Category/ModuleCat/Differentials/Basic.lean b/Mathlib/Algebra/Category/ModuleCat/Differentials/Basic.lean index 323f048f1af4a..cb4c1a9e98df5 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Differentials/Basic.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Differentials/Basic.lean @@ -112,12 +112,13 @@ noncomputable abbrev d (b : B) : KaehlerDifferential f := (D f).d b lemma ext {M : ModuleCat B} {α β : KaehlerDifferential f ⟶ M} (h : ∀ (b : B), α (d b) = β (d b)) : α = β := by rw [← sub_eq_zero] - have : ⊤ ≤ LinearMap.ker (α - β) := by + have : ⊤ ≤ LinearMap.ker (α - β).hom := by rw [← KaehlerDifferential.span_range_derivation, Submodule.span_le] rintro _ ⟨y, rfl⟩ - rw [SetLike.mem_coe, LinearMap.mem_ker, LinearMap.sub_apply, sub_eq_zero] + rw [SetLike.mem_coe, LinearMap.mem_ker, ModuleCat.hom_sub, LinearMap.sub_apply, sub_eq_zero] apply h rw [top_le_iff, LinearMap.ker_eq_top] at this + ext : 1 exact this /-- The map `KaehlerDifferential f ⟶ (ModuleCat.restrictScalars g').obj (KaehlerDifferential f')` @@ -133,6 +134,9 @@ noncomputable def map : letI := (g ≫ f').toAlgebra have : IsScalarTower A A' B' := IsScalarTower.of_algebraMap_eq' rfl have := IsScalarTower.of_algebraMap_eq' fac + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(Y := ...)`. + -- This suggests `restrictScalars` needs to be redesigned. + ModuleCat.ofHom (Y := (ModuleCat.restrictScalars g').obj (KaehlerDifferential f')) { toFun := fun x ↦ _root_.KaehlerDifferential.map A A' B B' x map_add' := by simp map_smul' := by simp } @@ -163,7 +167,7 @@ morphism `CommRingCat.KaehlerDifferential f ⟶ M`. -/ noncomputable def desc : CommRingCat.KaehlerDifferential f ⟶ M := letI := f.toAlgebra letI := Module.compHom M f - D.liftKaehlerDifferential + ofHom D.liftKaehlerDifferential @[simp] lemma desc_d (b : B) : D.desc (CommRingCat.KaehlerDifferential.d b) = D.d b := by diff --git a/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean b/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean index 5ea1401c768c2..9df2e8a46993e 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean @@ -77,7 +77,7 @@ variable (d : M.Derivation φ) /-- The postcomposition of a derivation by a morphism of presheaves of modules. -/ @[simps! d_apply] def postcomp (f : M ⟶ N) : N.Derivation φ where - d := (f.app _).toAddMonoidHom.comp d.d + d := (f.app _).hom.toAddMonoidHom.comp d.d d_map {X Y} g x := by simpa using naturality_apply f g (d.d x) d_app {X} a := by dsimp @@ -191,7 +191,7 @@ attribute [simp] relativeDifferentials'_obj lemma relativeDifferentials'_map_d {X Y : Dᵒᵖ} (f : X ⟶ Y) (x : R.obj X) : DFunLike.coe (α := CommRingCat.KaehlerDifferential (φ'.app X)) (β := fun _ ↦ CommRingCat.KaehlerDifferential (φ'.app Y)) - ((relativeDifferentials' φ').map f) (CommRingCat.KaehlerDifferential.d x) = + ((relativeDifferentials' φ').map f).hom (CommRingCat.KaehlerDifferential.d x) = CommRingCat.KaehlerDifferential.d (R.map f x) := CommRingCat.KaehlerDifferential.map_d (φ'.naturality f) _ diff --git a/Mathlib/Algebra/Category/ModuleCat/EpiMono.lean b/Mathlib/Algebra/Category/ModuleCat/EpiMono.lean index 4ba5ef78289de..c6db67923b764 100644 --- a/Mathlib/Algebra/Category/ModuleCat/EpiMono.lean +++ b/Mathlib/Algebra/Category/ModuleCat/EpiMono.lean @@ -24,22 +24,24 @@ namespace ModuleCat variable {R : Type u} [Ring R] {X Y : ModuleCat.{v} R} (f : X ⟶ Y) variable {M : Type v} [AddCommGroup M] [Module R M] -theorem ker_eq_bot_of_mono [Mono f] : LinearMap.ker f = ⊥ := - LinearMap.ker_eq_bot_of_cancel fun u v => (@cancel_mono _ _ _ _ _ f _ (↟u) (↟v)).1 +theorem ker_eq_bot_of_mono [Mono f] : LinearMap.ker f.hom = ⊥ := + LinearMap.ker_eq_bot_of_cancel fun u v h => ModuleCat.hom_ext_iff.mp <| + (@cancel_mono _ _ _ _ _ f _ (↟u) (↟v)).1 <| ModuleCat.hom_ext_iff.mpr h -theorem range_eq_top_of_epi [Epi f] : LinearMap.range f = ⊤ := - LinearMap.range_eq_top_of_cancel fun u v => (@cancel_epi _ _ _ _ _ f _ (↟u) (↟v)).1 +theorem range_eq_top_of_epi [Epi f] : LinearMap.range f.hom = ⊤ := + LinearMap.range_eq_top_of_cancel fun u v h => ModuleCat.hom_ext_iff.mp <| + (@cancel_epi _ _ _ _ _ f _ (↟u) (↟v)).1 <| ModuleCat.hom_ext_iff.mpr h -theorem mono_iff_ker_eq_bot : Mono f ↔ LinearMap.ker f = ⊥ := +theorem mono_iff_ker_eq_bot : Mono f ↔ LinearMap.ker f.hom = ⊥ := ⟨fun _ => ker_eq_bot_of_mono _, fun hf => ConcreteCategory.mono_of_injective _ <| by convert LinearMap.ker_eq_bot.1 hf⟩ theorem mono_iff_injective : Mono f ↔ Function.Injective f := by rw [mono_iff_ker_eq_bot, LinearMap.ker_eq_bot] -theorem epi_iff_range_eq_top : Epi f ↔ LinearMap.range f = ⊤ := +theorem epi_iff_range_eq_top : Epi f ↔ LinearMap.range f.hom = ⊤ := ⟨fun _ => range_eq_top_of_epi _, fun hf => - ConcreteCategory.epi_of_surjective _ <| LinearMap.range_eq_top.1 hf⟩ + ConcreteCategory.epi_of_surjective _ <| by convert LinearMap.range_eq_top.1 hf⟩ theorem epi_iff_surjective : Epi f ↔ Function.Surjective f := by rw [epi_iff_range_eq_top, LinearMap.range_eq_top] @@ -48,10 +50,10 @@ theorem epi_iff_surjective : Epi f ↔ Function.Surjective f := by def uniqueOfEpiZero (X) [h : Epi (0 : X ⟶ of R M)] : Unique M := uniqueOfSurjectiveZero X ((ModuleCat.epi_iff_surjective _).mp h) -instance mono_as_hom'_subtype (U : Submodule R X) : Mono (ModuleCat.asHomRight U.subtype) := +instance mono_as_hom'_subtype (U : Submodule R X) : Mono (ModuleCat.ofHom U.subtype) := (mono_iff_ker_eq_bot _).mpr (Submodule.ker_subtype U) -instance epi_as_hom''_mkQ (U : Submodule R X) : Epi (↿U.mkQ) := +instance epi_as_hom''_mkQ (U : Submodule R X) : Epi (ModuleCat.ofHom U.mkQ) := (epi_iff_range_eq_top _).mpr <| Submodule.range_mkQ _ instance forget_preservesEpimorphisms : (forget (ModuleCat.{v} R)).PreservesEpimorphisms where diff --git a/Mathlib/Algebra/Category/ModuleCat/FilteredColimits.lean b/Mathlib/Algebra/Category/ModuleCat/FilteredColimits.lean index 553ce84f4002e..08ca1ad384315 100644 --- a/Mathlib/Algebra/Category/ModuleCat/FilteredColimits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/FilteredColimits.lean @@ -132,7 +132,8 @@ def colimit : ModuleCatMax.{v, u, u} R := /-- The linear map from a given `R`-module in the diagram to the colimit module. -/ def coconeMorphism (j : J) : F.obj j ⟶ colimit F := - { (AddCommGrp.FilteredColimits.colimitCocone + ofHom + { (AddCommGrp.FilteredColimits.colimitCocone (F ⋙ forget₂ (ModuleCat R) AddCommGrp.{max v u})).ι.app j with map_smul' := fun r x => by erw [colimit_smul_mk_eq F r ⟨j, x⟩]; rfl } @@ -142,7 +143,7 @@ def colimitCocone : Cocone F where ι := { app := coconeMorphism F naturality := fun _ _' f => - LinearMap.coe_injective + hom_ext <| LinearMap.coe_injective ((Types.TypeMax.colimitCocone (F ⋙ forget (ModuleCat R))).ι.naturality f) } /-- Given a cocone `t` of `F`, the induced monoid linear map from the colimit to the cocone point. @@ -150,25 +151,27 @@ We already know that this is a morphism between additive groups. The only thing it is a linear map, i.e. preserves scalar multiplication. -/ def colimitDesc (t : Cocone F) : colimit F ⟶ t.pt := - { (AddCommGrp.FilteredColimits.colimitCoconeIsColimit + ofHom + { (AddCommGrp.FilteredColimits.colimitCoconeIsColimit (F ⋙ forget₂ (ModuleCatMax.{v, u} R) AddCommGrp.{max v u})).desc ((forget₂ (ModuleCat R) AddCommGrp.{max v u}).mapCocone t) with map_smul' := fun r x => by refine Quot.inductionOn x ?_; clear x; intro x; obtain ⟨j, x⟩ := x erw [colimit_smul_mk_eq] - exact LinearMap.map_smul (t.ι.app j) r x } + exact LinearMap.map_smul (t.ι.app j).hom r x } /-- The proposed colimit cocone is a colimit in `ModuleCat R`. -/ def colimitCoconeIsColimit : IsColimit (colimitCocone F) where desc := colimitDesc F fac t j := - LinearMap.coe_injective <| + hom_ext <| LinearMap.coe_injective <| (Types.TypeMax.colimitCoconeIsColimit.{v, u} (F ⋙ forget (ModuleCat R))).fac ((forget (ModuleCat R)).mapCocone t) j uniq t _ h := - LinearMap.coe_injective <| + hom_ext <| LinearMap.coe_injective <| (Types.TypeMax.colimitCoconeIsColimit (F ⋙ forget (ModuleCat R))).uniq - ((forget (ModuleCat R)).mapCocone t) _ fun j => funext fun x => LinearMap.congr_fun (h j) x + ((forget (ModuleCat R)).mapCocone t) _ fun j => funext fun x => LinearMap.congr_fun + (ModuleCat.hom_ext_iff.mp (h j)) x instance forget₂AddCommGroup_preservesFilteredColimits : PreservesFilteredColimits (forget₂ (ModuleCat.{u} R) AddCommGrp.{u}) where diff --git a/Mathlib/Algebra/Category/ModuleCat/Free.lean b/Mathlib/Algebra/Category/ModuleCat/Free.lean index 9f822325b9425..27fcdedc44ea4 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Free.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Free.lean @@ -48,7 +48,7 @@ theorem disjoint_span_sum : Disjoint (span R (range (u ∘ Sum.inl))) (span R (range (u ∘ Sum.inr))) := by rw [huv, disjoint_comm] refine Disjoint.mono_right (span_mono (range_comp_subset_range _ _)) ?_ - rw [← LinearMap.range_coe, span_eq (LinearMap.range S.f), hS.moduleCat_range_eq_ker] + rw [← LinearMap.range_coe, span_eq (LinearMap.range S.f.hom), hS.moduleCat_range_eq_ker] exact range_ker_disjoint hw include hv hm in @@ -66,8 +66,8 @@ where the top row is an exact sequence of modules and the maps on the bottom are independent. -/ theorem linearIndependent_leftExact : LinearIndependent R u := by rw [linearIndependent_sum] - refine ⟨?_, LinearIndependent.of_comp S.g hw, disjoint_span_sum hS hw huv⟩ - rw [huv, LinearMap.linearIndependent_iff S.f]; swap + refine ⟨?_, LinearIndependent.of_comp S.g.hom hw, disjoint_span_sum hS hw huv⟩ + rw [huv, LinearMap.linearIndependent_iff S.f.hom]; swap · rw [LinearMap.ker_eq_bot, ← mono_iff_injective] infer_instance exact hv @@ -78,7 +78,7 @@ include hS' hv in /-- Given a short exact sequence `0 ⟶ X₁ ⟶ X₂ ⟶ X₃ ⟶ 0` of `R`-modules and linearly independent families `v : ι → N` and `w : ι' → P`, we get a linearly independent family `ι ⊕ ι' → M` -/ theorem linearIndependent_shortExact {w : ι' → S.X₃} (hw : LinearIndependent R w) : - LinearIndependent R (Sum.elim (S.f ∘ v) (S.g.toFun.invFun ∘ w)) := by + LinearIndependent R (Sum.elim (S.f ∘ v) (S.g.hom.toFun.invFun ∘ w)) := by apply linearIndependent_leftExact hS'.exact hv _ hS'.mono_f rfl dsimp convert hw @@ -109,7 +109,7 @@ theorem span_exact {β : Type*} {u : ι ⊕ β → S.X₂} (huv : u ∘ Sum.inl rw [Finsupp.mem_span_range_iff_exists_finsupp] at hgm obtain ⟨cm, hm⟩ := hgm let m' : S.X₂ := Finsupp.sum cm fun j a ↦ a • (u (Sum.inr j)) - have hsub : m - m' ∈ LinearMap.range S.f := by + have hsub : m - m' ∈ LinearMap.range S.f.hom := by rw [hS.moduleCat_range_eq_ker] simp only [LinearMap.mem_ker, map_sub, sub_eq_zero] rw [← hm, map_finsupp_sum] @@ -138,7 +138,7 @@ include hS in families `v : ι → X₁` and `w : ι' → X₃`, we get a spanning family `ι ⊕ ι' → X₂` -/ theorem span_rightExact {w : ι' → S.X₃} (hv : ⊤ ≤ span R (range v)) (hw : ⊤ ≤ span R (range w)) (hE : Epi S.g) : - ⊤ ≤ span R (range (Sum.elim (S.f ∘ v) (S.g.toFun.invFun ∘ w))) := by + ⊤ ≤ span R (range (Sum.elim (S.f ∘ v) (S.g.hom.toFun.invFun ∘ w))) := by refine span_exact hS ?_ hv ?_ · simp only [AddHom.toFun_eq_coe, LinearMap.coe_toAddHom, Sum.elim_comp_inl] · convert hw diff --git a/Mathlib/Algebra/Category/ModuleCat/Images.lean b/Mathlib/Algebra/Category/ModuleCat/Images.lean index d37543fe17e12..17f0387b646f9 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Images.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Images.lean @@ -32,18 +32,18 @@ section -- implementation details of `HasImage` for ModuleCat; use the API, not these /-- The image of a morphism in `ModuleCat R` is just the bundling of `LinearMap.range f` -/ def image : ModuleCat R := - ModuleCat.of R (LinearMap.range f) + ModuleCat.of R (LinearMap.range f.hom) /-- The inclusion of `image f` into the target -/ def image.ι : image f ⟶ H := - f.range.subtype + ofHom f.hom.range.subtype instance : Mono (image.ι f) := ConcreteCategory.mono_of_injective (image.ι f) Subtype.val_injective /-- The corestriction map to the image -/ def factorThruImage : G ⟶ image f := - f.rangeRestrict + ofHom f.hom.rangeRestrict theorem image.fac : factorThruImage f ≫ image.ι f = f := rfl @@ -53,22 +53,23 @@ attribute [local simp] image.fac variable {f} /-- The universal property for the image factorisation -/ -noncomputable def image.lift (F' : MonoFactorisation f) : image f ⟶ F'.I where - toFun := (fun x => F'.e (Classical.indefiniteDescription _ x.2).1 : image f → F'.I) - map_add' x y := by - apply (mono_iff_injective F'.m).1 - · infer_instance - rw [LinearMap.map_add] - change (F'.e ≫ F'.m) _ = (F'.e ≫ F'.m) _ + (F'.e ≫ F'.m) _ - simp_rw [F'.fac, (Classical.indefiniteDescription (fun z => f z = _) _).2] - rfl - map_smul' c x := by - apply (mono_iff_injective F'.m).1 - · infer_instance - rw [LinearMap.map_smul] - change (F'.e ≫ F'.m) _ = _ • (F'.e ≫ F'.m) _ - simp_rw [F'.fac, (Classical.indefiniteDescription (fun z => f z = _) _).2] - rfl +noncomputable def image.lift (F' : MonoFactorisation f) : image f ⟶ F'.I := + ofHom + { toFun := (fun x => F'.e (Classical.indefiniteDescription _ x.2).1 : image f → F'.I) + map_add' := fun x y => by + apply (mono_iff_injective F'.m).1 + · infer_instance + rw [LinearMap.map_add] + change (F'.e ≫ F'.m) _ = (F'.e ≫ F'.m) _ + (F'.e ≫ F'.m) _ + simp_rw [F'.fac, (Classical.indefiniteDescription (fun z => f z = _) _).2] + rfl + map_smul' := fun c x => by + apply (mono_iff_injective F'.m).1 + · infer_instance + rw [LinearMap.map_smul] + change (F'.e ≫ F'.m) _ = _ • (F'.e ≫ F'.m) _ + simp_rw [F'.fac, (Classical.indefiniteDescription (fun z => f z = _) _).2] + rfl } theorem image.lift_fac (F' : MonoFactorisation f) : image.lift F' ≫ F'.m = image.ι f := by ext x @@ -92,17 +93,17 @@ noncomputable def isImage : IsImage (monoFactorisation f) where /-- The categorical image of a morphism in `ModuleCat R` agrees with the linear algebraic range. -/ noncomputable def imageIsoRange {G H : ModuleCat.{v} R} (f : G ⟶ H) : - Limits.image f ≅ ModuleCat.of R (LinearMap.range f) := + Limits.image f ≅ ModuleCat.of R (LinearMap.range f.hom) := IsImage.isoExt (Image.isImage f) (isImage f) @[simp, reassoc, elementwise] theorem imageIsoRange_inv_image_ι {G H : ModuleCat.{v} R} (f : G ⟶ H) : - (imageIsoRange f).inv ≫ Limits.image.ι f = ModuleCat.ofHom f.range.subtype := + (imageIsoRange f).inv ≫ Limits.image.ι f = ModuleCat.ofHom f.hom.range.subtype := IsImage.isoExt_inv_m _ _ @[simp, reassoc, elementwise] theorem imageIsoRange_hom_subtype {G H : ModuleCat.{v} R} (f : G ⟶ H) : - (imageIsoRange f).hom ≫ ModuleCat.ofHom f.range.subtype = Limits.image.ι f := by + (imageIsoRange f).hom ≫ ModuleCat.ofHom f.hom.range.subtype = Limits.image.ι f := by rw [← imageIsoRange_inv_image_ι f, Iso.hom_inv_id_assoc] end ModuleCat diff --git a/Mathlib/Algebra/Category/ModuleCat/Injective.lean b/Mathlib/Algebra/Category/ModuleCat/Injective.lean index 972d5bbec68c8..e11b9d7185a1f 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Injective.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Injective.lean @@ -20,16 +20,17 @@ namespace Module theorem injective_object_of_injective_module [inj : Injective R M] : CategoryTheory.Injective (ModuleCat.of R M) where factors g f m := - have ⟨l, h⟩ := inj.out f ((ModuleCat.mono_iff_injective f).mp m) g - ⟨l, LinearMap.ext h⟩ + have ⟨l, h⟩ := inj.out f.hom ((ModuleCat.mono_iff_injective f).mp m) g.hom + ⟨ModuleCat.ofHom l, by ext x; simpa using h x⟩ theorem injective_module_of_injective_object [inj : CategoryTheory.Injective <| ModuleCat.of R M] : Module.Injective R M where out X Y _ _ _ _ f hf g := by have : CategoryTheory.Mono (ModuleCat.ofHom f) := (ModuleCat.mono_iff_injective _).mpr hf - obtain ⟨l, rfl⟩ := inj.factors (ModuleCat.ofHom g) (ModuleCat.ofHom f) - exact ⟨l, fun _ ↦ rfl⟩ + obtain ⟨l, h⟩ := inj.factors (ModuleCat.ofHom g) (ModuleCat.ofHom f) + obtain rfl := ModuleCat.hom_ext_iff.mp h + exact ⟨l.hom, fun _ => rfl⟩ theorem injective_iff_injective_object : Module.Injective R M ↔ diff --git a/Mathlib/Algebra/Category/ModuleCat/Kernels.lean b/Mathlib/Algebra/Category/ModuleCat/Kernels.lean index b74fb6f74650b..6f2398e63fd4f 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Kernels.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Kernels.lean @@ -26,40 +26,38 @@ variable {M N : ModuleCat.{v} R} (f : M ⟶ N) /-- The kernel cone induced by the concrete kernel. -/ def kernelCone : KernelFork f := -- Porting note: previously proven by tidy - KernelFork.ofι (ofHom f.ker.subtype) <| by ext x; cases x; assumption + KernelFork.ofι (ofHom f.hom.ker.subtype) <| by ext x; cases x; assumption /-- The kernel of a linear map is a kernel in the categorical sense. -/ def kernelIsLimit : IsLimit (kernelCone f) := Fork.IsLimit.mk _ - (fun s => + (fun s => ofHom <| -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11036): broken dot notation on LinearMap.ker - LinearMap.codRestrict (LinearMap.ker f) (Fork.ι s) fun c => + LinearMap.codRestrict (LinearMap.ker f.hom) (Fork.ι s).hom fun c => LinearMap.mem_ker.2 <| by -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [← @Function.comp_apply _ _ _ f (Fork.ι s) c, ← coe_comp] - rw [Fork.condition, HasZeroMorphisms.comp_zero (Fork.ι s) N] + erw [← @Function.comp_apply _ _ _ f (Fork.ι s) c, ← LinearMap.coe_comp] + rw [← ModuleCat.hom_comp, Fork.condition, HasZeroMorphisms.comp_zero (Fork.ι s) N] rfl) - (fun _ => LinearMap.subtype_comp_codRestrict _ _ _) fun s m h => - LinearMap.ext fun x => Subtype.ext_iff_val.2 (by simp [← h]; rfl) + (fun _ => hom_ext <| LinearMap.subtype_comp_codRestrict _ _ _) fun s m h => + hom_ext <| LinearMap.ext fun x => Subtype.ext_iff_val.2 (by simp [← h]; rfl) /-- The cokernel cocone induced by the projection onto the quotient. -/ def cokernelCocone : CokernelCofork f := - CokernelCofork.ofπ (ofHom f.range.mkQ) <| LinearMap.range_mkQ_comp _ + CokernelCofork.ofπ (ofHom f.hom.range.mkQ) <| hom_ext <| LinearMap.range_mkQ_comp _ /-- The projection onto the quotient is a cokernel in the categorical sense. -/ def cokernelIsColimit : IsColimit (cokernelCocone f) := Cofork.IsColimit.mk _ - (fun s => - f.range.liftQ (Cofork.π s) <| LinearMap.range_le_ker_iff.2 <| CokernelCofork.condition s) - (fun s => f.range.liftQ_mkQ (Cofork.π s) _) fun s m h => by + (fun s => ofHom <| f.hom.range.liftQ (Cofork.π s).hom <| + LinearMap.range_le_ker_iff.2 <| ModuleCat.hom_ext_iff.mp <| CokernelCofork.condition s) + (fun s => hom_ext <| f.hom.range.liftQ_mkQ (Cofork.π s).hom _) fun s m h => by -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11036): broken dot notation - haveI : Epi (ofHom (LinearMap.range f).mkQ) := + haveI : Epi (ofHom (LinearMap.range f.hom).mkQ) := (epi_iff_range_eq_top _).mpr (Submodule.range_mkQ _) -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11036): broken dot notation - apply (cancel_epi (ofHom (LinearMap.range f).mkQ)).1 - convert h - -- Porting note: no longer necessary - -- exact Submodule.liftQ_mkQ _ _ _ + apply (cancel_epi (ofHom (LinearMap.range f.hom).mkQ)).1 + exact h end @@ -84,20 +82,20 @@ agrees with the usual module-theoretical kernel. -/ noncomputable def kernelIsoKer {G H : ModuleCat.{v} R} (f : G ⟶ H) : -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11036): broken dot notation - kernel f ≅ ModuleCat.of R (LinearMap.ker f) := + kernel f ≅ ModuleCat.of R (LinearMap.ker f.hom) := limit.isoLimitCone ⟨_, kernelIsLimit f⟩ -- We now show this isomorphism commutes with the inclusion of the kernel into the source. @[simp, elementwise] -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11036): broken dot notation theorem kernelIsoKer_inv_kernel_ι : (kernelIsoKer f).inv ≫ kernel.ι f = - (LinearMap.ker f).subtype := + ofHom (LinearMap.ker f.hom).subtype := limit.isoLimitCone_inv_π _ _ @[simp, elementwise] theorem kernelIsoKer_hom_ker_subtype : -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11036): broken dot notation - (kernelIsoKer f).hom ≫ (LinearMap.ker f).subtype = kernel.ι f := + (kernelIsoKer f).hom ≫ ofHom (LinearMap.ker f.hom).subtype = kernel.ι f := IsLimit.conePointUniqueUpToIso_inv_comp _ (limit.isLimit _) WalkingParallelPair.zero /-- The categorical cokernel of a morphism in `ModuleCat` @@ -105,18 +103,18 @@ agrees with the usual module-theoretical quotient. -/ noncomputable def cokernelIsoRangeQuotient {G H : ModuleCat.{v} R} (f : G ⟶ H) : -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11036): broken dot notation - cokernel f ≅ ModuleCat.of R (H ⧸ LinearMap.range f) := + cokernel f ≅ ModuleCat.of R (H ⧸ LinearMap.range f.hom) := colimit.isoColimitCocone ⟨_, cokernelIsColimit f⟩ -- We now show this isomorphism commutes with the projection of target to the cokernel. @[simp, elementwise] theorem cokernel_π_cokernelIsoRangeQuotient_hom : - cokernel.π f ≫ (cokernelIsoRangeQuotient f).hom = f.range.mkQ := + cokernel.π f ≫ (cokernelIsoRangeQuotient f).hom = ofHom f.hom.range.mkQ := colimit.isoColimitCocone_ι_hom _ _ @[simp, elementwise] theorem range_mkQ_cokernelIsoRangeQuotient_inv : - ↿f.range.mkQ ≫ (cokernelIsoRangeQuotient f).inv = cokernel.π f := + ofHom f.hom.range.mkQ ≫ (cokernelIsoRangeQuotient f).inv = cokernel.π f := colimit.isoColimitCocone_ι_inv ⟨_, cokernelIsColimit f⟩ WalkingParallelPair.one theorem cokernel_π_ext {M N : ModuleCat.{u} R} (f : M ⟶ N) {x y : N} (m : M) (w : x = y + f m) : diff --git a/Mathlib/Algebra/Category/ModuleCat/Limits.lean b/Mathlib/Algebra/Category/ModuleCat/Limits.lean index 862547dd00ba8..607ca41120826 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Limits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Limits.lean @@ -45,9 +45,7 @@ def sectionsSubmodule : Submodule R (∀ j, F.obj j) := forget₂ AddCommGrp AddGrp.{w}) with carrier := (F ⋙ forget (ModuleCat R)).sections smul_mem' := fun r s sh j j' f => by - simp only [forget_map, Functor.comp_map, Pi.smul_apply, map_smul] - dsimp [Functor.sections] at sh - rw [sh f] } + simpa [Functor.sections, forget_map] using congr_arg (r • ·) (sh f) } instance : AddCommMonoid (F ⋙ forget (ModuleCat R)).sections := inferInstanceAs <| AddCommMonoid (sectionsSubmodule F) @@ -101,16 +99,16 @@ namespace HasLimits def limitCone : Cone F where pt := ModuleCat.of R (Types.Small.limitCone.{v, w} (F ⋙ forget _)).pt π := - { app := limitπLinearMap F - naturality := fun _ _ f => - LinearMap.coe_injective ((Types.Small.limitCone (F ⋙ forget _)).π.naturality f) } + { app := fun j => ofHom (limitπLinearMap F j) + naturality := fun _ _ f => hom_ext <| LinearMap.coe_injective <| + ((Types.Small.limitCone (F ⋙ forget _)).π.naturality f) } /-- Witness that the limit cone in `ModuleCat R` is a limit cone. (Internal use only; use the limits API.) -/ def limitConeIsLimit : IsLimit (limitCone.{t, v, w} F) := by refine IsLimit.ofFaithful (forget (ModuleCat R)) (Types.Small.limitConeIsLimit.{v, w} _) - (fun s => ⟨⟨(Types.Small.limitConeIsLimit.{v, w} _).lift + (fun s => ofHom ⟨⟨(Types.Small.limitConeIsLimit.{v, w} _).lift ((forget (ModuleCat R)).mapCone s), ?_⟩, ?_⟩) (fun s => rfl) · intro x y @@ -216,14 +214,12 @@ variable (f : ∀ i j, i ≤ j → G i →ₗ[R] G j) [DirectedSystem G fun i j @[simps] def directLimitDiagram : ι ⥤ ModuleCat R where obj i := ModuleCat.of R (G i) - map hij := f _ _ hij.le + map hij := ofHom (f _ _ hij.le) map_id i := by - apply LinearMap.ext - intro x + ext apply Module.DirectedSystem.map_self map_comp hij hjk := by - apply LinearMap.ext - intro x + ext symm apply Module.DirectedSystem.map_map @@ -237,35 +233,33 @@ In `directLimitIsColimit` we show that it is a colimit cocone. -/ def directLimitCocone : Cocone (directLimitDiagram G f) where pt := ModuleCat.of R <| DirectLimit G f ι := - { app := Module.DirectLimit.of R ι G f + { app := fun x => ofHom (Module.DirectLimit.of R ι G f x) naturality := fun _ _ hij => by - apply LinearMap.ext - intro x + ext exact DirectLimit.of_f } /-- The unbundled `directLimit` of modules is a colimit in the sense of `CategoryTheory`. -/ @[simps] def directLimitIsColimit [IsDirected ι (· ≤ ·)] : IsColimit (directLimitCocone G f) where - desc s := - DirectLimit.lift R ι G f s.ι.app fun i j h x => by + desc s := ofHom <| + DirectLimit.lift R ι G f (fun i => (s.ι.app i).hom) fun i j h x => by + simp only [Functor.const_obj_obj] rw [← s.w (homOfLE h)] rfl fac s i := by - apply LinearMap.ext - intro x + ext dsimp only [directLimitCocone, CategoryStruct.comp] rw [LinearMap.comp_apply] apply DirectLimit.lift_of uniq s m h := by have : s.ι.app = fun i => - LinearMap.comp m (DirectLimit.of R ι (fun i => G i) (fun i j H => f i j H) i) := by + (ofHom (DirectLimit.of R ι (fun i => G i) (fun i j H => f i j H) i)) ≫ m := by funext i rw [← h] rfl - apply LinearMap.ext - intro x + ext simp only [this] apply Module.DirectLimit.lift_unique diff --git a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean index d6e565553b5e4..89bdba339ff9f 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean @@ -54,25 +54,27 @@ def tensorObj (M N : ModuleCat R) : ModuleCat R := /-- (implementation) tensor product of morphisms R-modules -/ def tensorHom {M N M' N' : ModuleCat R} (f : M ⟶ N) (g : M' ⟶ N') : tensorObj M M' ⟶ tensorObj N N' := - TensorProduct.map f g + ofHom <| TensorProduct.map f.hom g.hom /-- (implementation) left whiskering for R-modules -/ def whiskerLeft (M : ModuleCat R) {N₁ N₂ : ModuleCat R} (f : N₁ ⟶ N₂) : tensorObj M N₁ ⟶ tensorObj M N₂ := - f.lTensor M + ofHom <| f.hom.lTensor M /-- (implementation) right whiskering for R-modules -/ def whiskerRight {M₁ M₂ : ModuleCat R} (f : M₁ ⟶ M₂) (N : ModuleCat R) : tensorObj M₁ N ⟶ tensorObj M₂ N := - f.rTensor N + ofHom <| f.hom.rTensor N theorem tensor_id (M N : ModuleCat R) : tensorHom (𝟙 M) (𝟙 N) = 𝟙 (ModuleCat.of R (M ⊗ N)) := by + ext : 1 -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11041): even with high priority `ext` fails to find this. apply TensorProduct.ext rfl theorem tensor_comp {X₁ Y₁ Z₁ X₂ Y₂ Z₂ : ModuleCat R} (f₁ : X₁ ⟶ Y₁) (f₂ : X₂ ⟶ Y₂) (g₁ : Y₁ ⟶ Z₁) (g₂ : Y₂ ⟶ Z₂) : tensorHom (f₁ ≫ g₁) (f₂ ≫ g₂) = tensorHom f₁ f₂ ≫ tensorHom g₁ g₂ := by + ext : 1 -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11041): even with high priority `ext` fails to find this. apply TensorProduct.ext rfl @@ -95,63 +97,36 @@ instance instMonoidalCategoryStruct : MonoidalCategoryStruct (ModuleCat.{u} R) w tensorObj := tensorObj whiskerLeft := whiskerLeft whiskerRight := whiskerRight - tensorHom f g := TensorProduct.map f g + tensorHom f g := ofHom <| TensorProduct.map f.hom g.hom tensorUnit := ModuleCat.of R R associator := associator leftUnitor := leftUnitor rightUnitor := rightUnitor -section - -/-! The `associator_naturality` and `pentagon` lemmas below are very slow to elaborate. - -We give them some help by expressing the lemmas first non-categorically, then using -`convert _aux using 1` to have the elaborator work as little as possible. -/ - - -open TensorProduct (assoc map) - -private theorem associator_naturality_aux {X₁ X₂ X₃ : Type*} [AddCommMonoid X₁] [AddCommMonoid X₂] - [AddCommMonoid X₃] [Module R X₁] [Module R X₂] [Module R X₃] {Y₁ Y₂ Y₃ : Type*} - [AddCommMonoid Y₁] [AddCommMonoid Y₂] [AddCommMonoid Y₃] [Module R Y₁] [Module R Y₂] - [Module R Y₃] (f₁ : X₁ →ₗ[R] Y₁) (f₂ : X₂ →ₗ[R] Y₂) (f₃ : X₃ →ₗ[R] Y₃) : - ↑(assoc R Y₁ Y₂ Y₃) ∘ₗ map (map f₁ f₂) f₃ = map f₁ (map f₂ f₃) ∘ₗ ↑(assoc R X₁ X₂ X₃) := by - apply TensorProduct.ext_threefold - intro x y z - rfl - -variable (R) - -private theorem pentagon_aux (W X Y Z : Type*) [AddCommMonoid W] [AddCommMonoid X] - [AddCommMonoid Y] [AddCommMonoid Z] [Module R W] [Module R X] [Module R Y] [Module R Z] : - (((assoc R X Y Z).toLinearMap.lTensor W).comp - (assoc R W (X ⊗[R] Y) Z).toLinearMap).comp - ((assoc R W X Y).toLinearMap.rTensor Z) = - (assoc R W X (Y ⊗[R] Z)).toLinearMap.comp (assoc R (W ⊗[R] X) Y Z).toLinearMap := by - apply TensorProduct.ext_fourfold - intro w x y z - rfl - -end - theorem associator_naturality {X₁ X₂ X₃ Y₁ Y₂ Y₃ : ModuleCat R} (f₁ : X₁ ⟶ Y₁) (f₂ : X₂ ⟶ Y₂) (f₃ : X₃ ⟶ Y₃) : tensorHom (tensorHom f₁ f₂) f₃ ≫ (associator Y₁ Y₂ Y₃).hom = (associator X₁ X₂ X₃).hom ≫ tensorHom f₁ (tensorHom f₂ f₃) := by - convert associator_naturality_aux f₁ f₂ f₃ using 1 + ext : 1 + apply TensorProduct.ext_threefold + intro x y z + rfl theorem pentagon (W X Y Z : ModuleCat R) : whiskerRight (associator W X Y).hom Z ≫ (associator W (tensorObj X Y) Z).hom ≫ whiskerLeft W (associator X Y Z).hom = (associator (tensorObj W X) Y Z).hom ≫ (associator W X (tensorObj Y Z)).hom := by - convert pentagon_aux R W X Y Z using 1 + ext : 1 + apply TensorProduct.ext_fourfold + intro w x y z + rfl theorem leftUnitor_naturality {M N : ModuleCat R} (f : M ⟶ N) : tensorHom (𝟙 (ModuleCat.of R R)) f ≫ (leftUnitor N).hom = (leftUnitor M).hom ≫ f := by + ext : 1 -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11041): broken ext apply TensorProduct.ext - apply LinearMap.ext_ring - apply LinearMap.ext; intro x + ext x -- Porting note (https://github.com/leanprover-community/mathlib4/pull/10934): used to be dsimp change ((leftUnitor N).hom) ((tensorHom (𝟙 (of R R)) f) ((1 : R) ⊗ₜ[R] x)) = f (((leftUnitor M).hom) (1 ⊗ₜ[R] x)) @@ -161,13 +136,11 @@ theorem leftUnitor_naturality {M N : ModuleCat R} (f : M ⟶ N) : theorem rightUnitor_naturality {M N : ModuleCat R} (f : M ⟶ N) : tensorHom f (𝟙 (ModuleCat.of R R)) ≫ (rightUnitor N).hom = (rightUnitor M).hom ≫ f := by + ext : 1 -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11041): broken ext apply TensorProduct.ext - apply LinearMap.ext; intro x - apply LinearMap.ext_ring - -- Porting note (https://github.com/leanprover-community/mathlib4/pull/10934): used to be dsimp - change ((rightUnitor N).hom) ((tensorHom f (𝟙 (of R R))) (x ⊗ₜ[R] (1 : R))) = - f (((rightUnitor M).hom) (x ⊗ₜ[R] 1)) + ext x + dsimp erw [TensorProduct.rid_tmul, TensorProduct.rid_tmul] rw [LinearMap.map_smul] rfl @@ -175,9 +148,9 @@ theorem rightUnitor_naturality {M N : ModuleCat R} (f : M ⟶ N) : theorem triangle (M N : ModuleCat.{u} R) : (associator M (ModuleCat.of R R) N).hom ≫ tensorHom (𝟙 M) (leftUnitor N).hom = tensorHom (rightUnitor M).hom (𝟙 N) := by + ext : 1 apply TensorProduct.ext_threefold intro x y z - change R at y -- Porting note (https://github.com/leanprover-community/mathlib4/pull/10934): used to be dsimp [tensorHom, associator] change x ⊗ₜ[R] ((leftUnitor N).hom) (y ⊗ₜ[R] z) = ((rightUnitor M).hom) (x ⊗ₜ[R] y) ⊗ₜ[R] z erw [TensorProduct.lid_tmul, TensorProduct.rid_tmul] @@ -263,7 +236,7 @@ variable (f : M₁ → M₂ → M₃) (h₁ : ∀ m₁ m₂ n, f (m₁ + m₂) n /-- Construct for morphisms from the tensor product of two objects in `ModuleCat`. -/ noncomputable def tensorLift : M₁ ⊗ M₂ ⟶ M₃ := - TensorProduct.lift (LinearMap.mk₂ R f h₁ h₂ h₃ h₄) + ofHom <| TensorProduct.lift (LinearMap.mk₂ R f h₁ h₂ h₃ h₄) @[simp] lemma tensorLift_tmul (m : M₁) (n : M₂) : @@ -273,13 +246,13 @@ end lemma tensor_ext {f g : M₁ ⊗ M₂ ⟶ M₃} (h : ∀ m n, f (m ⊗ₜ n) = g (m ⊗ₜ n)) : f = g := - TensorProduct.ext (by ext; apply h) + hom_ext <| TensorProduct.ext (by ext; apply h) /-- Extensionality lemma for morphisms from a module of the form `(M₁ ⊗ M₂) ⊗ M₃`. -/ lemma tensor_ext₃' {f g : (M₁ ⊗ M₂) ⊗ M₃ ⟶ M₄} (h : ∀ m₁ m₂ m₃, f (m₁ ⊗ₜ m₂ ⊗ₜ m₃) = g (m₁ ⊗ₜ m₂ ⊗ₜ m₃)) : f = g := - TensorProduct.ext_threefold h + hom_ext <| TensorProduct.ext_threefold h /-- Extensionality lemma for morphisms from a module of the form `M₁ ⊗ (M₂ ⊗ M₃)`. -/ lemma tensor_ext₃ {f g : M₁ ⊗ (M₂ ⊗ M₃) ⟶ M₄} @@ -292,56 +265,57 @@ end MonoidalCategory open Opposite --- Porting note: simp wasn't firing but rw was, annoying instance : MonoidalPreadditive (ModuleCat.{u} R) := by refine ⟨?_, ?_, ?_, ?_⟩ · intros + ext : 1 refine TensorProduct.ext (LinearMap.ext fun x => LinearMap.ext fun y => ?_) - simp only [LinearMap.compr₂_apply, TensorProduct.mk_apply] - rw [LinearMap.zero_apply] + simp only [LinearMap.compr₂_apply, TensorProduct.mk_apply, hom_zero, LinearMap.zero_apply] -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 erw [MonoidalCategory.whiskerLeft_apply] - rw [LinearMap.zero_apply, TensorProduct.tmul_zero] + simp · intros + ext : 1 refine TensorProduct.ext (LinearMap.ext fun x => LinearMap.ext fun y => ?_) - simp only [LinearMap.compr₂_apply, TensorProduct.mk_apply] - rw [LinearMap.zero_apply] + simp only [LinearMap.compr₂_apply, TensorProduct.mk_apply, hom_zero, LinearMap.zero_apply, ] -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 erw [MonoidalCategory.whiskerRight_apply] - rw [LinearMap.zero_apply, TensorProduct.zero_tmul] + simp · intros + ext : 1 refine TensorProduct.ext (LinearMap.ext fun x => LinearMap.ext fun y => ?_) - simp only [LinearMap.compr₂_apply, TensorProduct.mk_apply] - rw [LinearMap.add_apply] + simp only [LinearMap.compr₂_apply, TensorProduct.mk_apply, hom_add, LinearMap.add_apply] -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 erw [MonoidalCategory.whiskerLeft_apply, MonoidalCategory.whiskerLeft_apply] erw [MonoidalCategory.whiskerLeft_apply] - rw [LinearMap.add_apply, TensorProduct.tmul_add] + simp [TensorProduct.tmul_add] · intros + ext : 1 refine TensorProduct.ext (LinearMap.ext fun x => LinearMap.ext fun y => ?_) - simp only [LinearMap.compr₂_apply, TensorProduct.mk_apply] - rw [LinearMap.add_apply] + simp only [LinearMap.compr₂_apply, TensorProduct.mk_apply, hom_add, LinearMap.add_apply] -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 erw [MonoidalCategory.whiskerRight_apply, MonoidalCategory.whiskerRight_apply] erw [MonoidalCategory.whiskerRight_apply] - rw [LinearMap.add_apply, TensorProduct.add_tmul] + simp [TensorProduct.add_tmul] --- Porting note: simp wasn't firing but rw was, annoying instance : MonoidalLinear R (ModuleCat.{u} R) := by refine ⟨?_, ?_⟩ · intros + ext : 1 refine TensorProduct.ext (LinearMap.ext fun x => LinearMap.ext fun y => ?_) - simp only [LinearMap.compr₂_apply, TensorProduct.mk_apply] - rw [LinearMap.smul_apply] + simp only [LinearMap.compr₂_apply, TensorProduct.mk_apply, hom_smul, LinearMap.smul_apply] -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 erw [MonoidalCategory.whiskerLeft_apply, MonoidalCategory.whiskerLeft_apply] - rw [LinearMap.smul_apply, TensorProduct.tmul_smul] + simp · intros + ext : 1 refine TensorProduct.ext (LinearMap.ext fun x => LinearMap.ext fun y => ?_) - simp only [LinearMap.compr₂_apply, TensorProduct.mk_apply] - rw [LinearMap.smul_apply] + simp only [LinearMap.compr₂_apply, TensorProduct.mk_apply, hom_smul, LinearMap.smul_apply] -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 erw [MonoidalCategory.whiskerRight_apply, MonoidalCategory.whiskerRight_apply] - rw [LinearMap.smul_apply, TensorProduct.smul_tmul, TensorProduct.tmul_smul] + simp [TensorProduct.smul_tmul, TensorProduct.tmul_smul] + +@[simp] lemma ofHom₂_compr₂ {M N P Q : ModuleCat.{u} R} (f : M →ₗ[R] N →ₗ[R] P) (g : P →ₗ[R] Q): + ofHom₂ (f.compr₂ g) = ofHom₂ f ≫ ofHom (Linear.rightComp R _ (ofHom g)) := rfl end ModuleCat diff --git a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Closed.lean b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Closed.lean index 3b8592dcc5ce7..61228a85b17e0 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Closed.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Closed.lean @@ -27,15 +27,13 @@ variable {R : Type u} [CommRing R] def monoidalClosedHomEquiv (M N P : ModuleCat.{u} R) : ((MonoidalCategory.tensorLeft M).obj N ⟶ P) ≃ (N ⟶ ((linearCoyoneda R (ModuleCat R)).obj (op M)).obj P) where - toFun f := LinearMap.compr₂ (TensorProduct.mk R N M) ((β_ N M).hom ≫ f) - invFun f := (β_ M N).hom ≫ TensorProduct.lift f + toFun f := ofHom₂ <| LinearMap.compr₂ (TensorProduct.mk R N M) ((β_ N M).hom ≫ f).hom + invFun f := (β_ M N).hom ≫ ofHom (TensorProduct.lift f.hom₂) left_inv f := by + ext : 1 apply TensorProduct.ext' intro m n - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [coe_comp] - rw [Function.comp_apply] - -- This used to be `rw` and was longer (?), but we need `erw` after https://github.com/leanprover/lean4/pull/2644 + simp only [Hom.hom₂_ofHom₂, LinearMap.comp_apply, hom_comp, MonoidalCategory.tensorLeft_obj] erw [MonoidalCategory.braiding_hom_apply, TensorProduct.lift.tmul] right_inv _ := rfl @@ -47,6 +45,7 @@ instance : MonoidalClosed (ModuleCat.{u} R) where -- Porting note: this proof was automatic in mathlib3 homEquiv_naturality_left_symm := by intros + ext : 1 apply TensorProduct.ext' intro m n rfl } } @@ -57,26 +56,24 @@ theorem ihom_map_apply {M N P : ModuleCat.{u} R} (f : N ⟶ P) (g : ModuleCat.of open MonoidalCategory --- Porting note: `CoeFun` was replaced by `DFunLike` --- I can't seem to express the function coercion here without writing `@DFunLike.coe`. theorem monoidalClosed_curry {M N P : ModuleCat.{u} R} (f : M ⊗ N ⟶ P) (x : M) (y : N) : - @DFunLike.coe _ _ _ LinearMap.instFunLike - ((MonoidalClosed.curry f : N →ₗ[R] M →ₗ[R] P) y) x = f (x ⊗ₜ[R] y) := + ((MonoidalClosed.curry f).hom y).hom x = f (x ⊗ₜ[R] y) := rfl @[simp] theorem monoidalClosed_uncurry {M N P : ModuleCat.{u} R} (f : N ⟶ M ⟶[ModuleCat.{u} R] P) (x : M) (y : N) : - MonoidalClosed.uncurry f (x ⊗ₜ[R] y) = - @DFunLike.coe _ _ _ LinearMap.instFunLike (f y) x := + MonoidalClosed.uncurry f (x ⊗ₜ[R] y) = (f y).hom x := rfl /-- Describes the counit of the adjunction `M ⊗ - ⊣ Hom(M, -)`. Given an `R`-module `N` this should give a map `M ⊗ Hom(M, N) ⟶ N`, so we flip the order of the arguments in the identity map `Hom(M, N) ⟶ (M ⟶ N)` and uncurry the resulting map `M ⟶ Hom(M, N) ⟶ N.` -/ theorem ihom_ev_app (M N : ModuleCat.{u} R) : - (ihom.ev M).app N = TensorProduct.uncurry _ _ _ _ LinearMap.id.flip := by + (ihom.ev M).app N = ModuleCat.ofHom (TensorProduct.uncurry R M ((ihom M).obj N) N + (LinearMap.lcomp _ _ homLinearEquiv.toLinearMap ∘ₗ LinearMap.id.flip)) := by rw [← MonoidalClosed.uncurry_id_eq_ev] + ext : 1 apply TensorProduct.ext' apply monoidalClosed_uncurry @@ -84,11 +81,12 @@ theorem ihom_ev_app (M N : ModuleCat.{u} R) : define a map `N ⟶ Hom(M, M ⊗ N)`, which is given by flipping the arguments in the natural `R`-bilinear map `M ⟶ N ⟶ M ⊗ N`. -/ theorem ihom_coev_app (M N : ModuleCat.{u} R) : - (ihom.coev M).app N = (TensorProduct.mk _ _ _).flip := + (ihom.coev M).app N = ModuleCat.ofHom₂ (TensorProduct.mk _ _ _).flip := rfl theorem monoidalClosed_pre_app {M N : ModuleCat.{u} R} (P : ModuleCat.{u} R) (f : N ⟶ M) : - (MonoidalClosed.pre f).app P = LinearMap.lcomp R _ f := + (MonoidalClosed.pre f).app P = ofHom (homLinearEquiv.symm.toLinearMap ∘ₗ + LinearMap.lcomp _ _ f.hom ∘ₗ homLinearEquiv.toLinearMap) := rfl end ModuleCat diff --git a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Symmetric.lean b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Symmetric.lean index d6b5c46f9dba1..42370456d2060 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Symmetric.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Symmetric.lean @@ -29,6 +29,7 @@ namespace MonoidalCategory @[simp] theorem braiding_naturality {X₁ X₂ Y₁ Y₂ : ModuleCat.{u} R} (f : X₁ ⟶ Y₁) (g : X₂ ⟶ Y₂) : (f ⊗ g) ≫ (Y₁.braiding Y₂).hom = (X₁.braiding X₂).hom ≫ (g ⊗ f) := by + ext : 1 apply TensorProduct.ext' intro x y rfl @@ -49,6 +50,7 @@ theorem braiding_naturality_right (X : ModuleCat R) {Y Z : ModuleCat R} (f : Y theorem hexagon_forward (X Y Z : ModuleCat.{u} R) : (α_ X Y Z).hom ≫ (braiding X _).hom ≫ (α_ Y Z X).hom = (braiding X Y).hom ▷ Z ≫ (α_ Y X Z).hom ≫ Y ◁ (braiding X Z).hom := by + ext : 1 apply TensorProduct.ext_threefold intro x y z rfl @@ -58,6 +60,7 @@ theorem hexagon_reverse (X Y Z : ModuleCat.{u} R) : (α_ X Y Z).inv ≫ (braiding _ Z).hom ≫ (α_ Z X Y).inv = X ◁ (Y.braiding Z).hom ≫ (α_ X Z Y).inv ≫ (X.braiding Z).hom ▷ Y := by apply (cancel_epi (α_ X Y Z).hom).1 + ext : 1 apply TensorProduct.ext_threefold intro x y z rfl @@ -74,6 +77,7 @@ instance symmetricCategory : SymmetricCategory (ModuleCat.{u} R) where -- Porting note: this proof was automatic in Lean3 -- now `aesop` is applying `ModuleCat.ext` in favour of `TensorProduct.ext`. symmetry _ _ := by + ext : 1 apply TensorProduct.ext' aesop_cat @@ -88,8 +92,8 @@ theorem braiding_inv_apply {M N : ModuleCat.{u} R} (m : M) (n : N) : rfl theorem tensorμ_eq_tensorTensorTensorComm {A B C D : ModuleCat R} : - tensorμ A B C D = (TensorProduct.tensorTensorTensorComm R A B C D).toLinearMap := - TensorProduct.ext <| TensorProduct.ext <| LinearMap.ext₂ fun _ _ => + tensorμ A B C D = ofHom (TensorProduct.tensorTensorTensorComm R A B C D).toLinearMap := + ModuleCat.hom_ext <| TensorProduct.ext <| TensorProduct.ext <| LinearMap.ext₂ fun _ _ => TensorProduct.ext <| LinearMap.ext₂ fun _ _ => rfl @[simp] diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean index 653280c6b57bd..8755d8d2ac2f7 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean @@ -161,7 +161,10 @@ when the preferred constructor `PresheafOfModules.mk` is not as convenient as th @[simps] def ofPresheaf : PresheafOfModules.{v} R where obj X := ModuleCat.of _ (M.obj X) - map f := + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(Y := ...)`. + -- This suggests `restrictScalars` needs to be redesigned. + map {X Y} f := ModuleCat.ofHom + (Y := (ModuleCat.restrictScalars (R.map f)).obj (ModuleCat.of _ (M.obj Y))) { toFun := fun x ↦ M.map f x map_add' := by simp map_smul' := fun r m ↦ map_smul f r m } @@ -178,7 +181,7 @@ which satisfy a suitable linearity condition. -/ def homMk (φ : M₁.presheaf ⟶ M₂.presheaf) (hφ : ∀ (X : Cᵒᵖ) (r : R.obj X) (m : M₁.obj X), φ.app X (r • m) = r • φ.app X m) : M₁ ⟶ M₂ where - app X := + app X := ModuleCat.ofHom { toFun := φ.app X map_add' := by simp map_smul' := hφ X } @@ -197,30 +200,21 @@ instance : Neg (M₁ ⟶ M₂) where { app := fun X ↦ -f.app X naturality := fun {X Y} h ↦ by ext x - dsimp - erw [map_neg] - rw [← naturality_apply] - rfl } + simp [← naturality_apply] } instance : Add (M₁ ⟶ M₂) where add f g := { app := fun X ↦ f.app X + g.app X naturality := fun {X Y} h ↦ by ext x - dsimp - erw [map_add] - rw [← naturality_apply, ← naturality_apply] - rfl } + simp [← naturality_apply] } instance : Sub (M₁ ⟶ M₂) where sub f g := { app := fun X ↦ f.app X - g.app X naturality := fun {X Y} h ↦ by ext x - dsimp - erw [map_sub] - rw [← naturality_apply, ← naturality_apply] - rfl } + simp [← naturality_apply] } @[simp] lemma neg_app (f : M₁ ⟶ M₂) (X : Cᵒᵖ) : (-f).app X = -f.app X := rfl @[simp] lemma add_app (f g : M₁ ⟶ M₂) (X : Cᵒᵖ) : (f + g).app X = f.app X + g.app X := rfl @@ -267,7 +261,10 @@ noncomputable def restriction {X Y : Cᵒᵖ} (f : X ⟶ Y) : /-- The obvious free presheaf of modules of rank `1`. -/ def unit : PresheafOfModules R where obj X := ModuleCat.of _ (R.obj X) - map {X Y } f := + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(Y := ...)`. + -- This suggests `restrictScalars` needs to be redesigned. + map {X Y} f := ModuleCat.ofHom + (Y := (ModuleCat.restrictScalars (R.map f)).obj (ModuleCat.of (R.obj Y) (R.obj Y))) { toFun := fun x ↦ R.map f x map_add' := by simp map_smul' := by aesop_cat } @@ -321,14 +318,16 @@ def unitHomEquiv (M : PresheafOfModules R) : toFun f := sectionsMk (fun X ↦ Hom.app f X (1 : R.obj X)) (by intros; rw [← naturality_apply, unit_map_one]) invFun s := - { app := fun X ↦ (LinearMap.ringLmapEquivSelf (R.obj X) ℤ (M.obj X)).symm (s.val X) + { app := fun X ↦ ModuleCat.ofHom + ((LinearMap.ringLmapEquivSelf (R.obj X) ℤ (M.obj X)).symm (s.val X)) naturality := fun {X Y} f ↦ by - ext (x : R.obj X) - change R.map f x • s.eval Y = M.map f (x • s.eval X) + ext + dsimp + change R.map f 1 • s.eval Y = M.map f (1 • s.eval X) simp } left_inv f := by - ext1 X - exact (LinearMap.ringLmapEquivSelf (R.obj X) ℤ (M.obj X)).symm_apply_apply (f.app X) + ext X : 2 + exact (LinearMap.ringLmapEquivSelf (R.obj X) ℤ (M.obj X)).symm_apply_apply (f.app X).hom right_inv s := by ext X exact (LinearMap.ringLmapEquivSelf (R.obj X) ℤ (M.obj X)).apply_symm_apply (s.val X) @@ -352,27 +351,29 @@ variable (M : PresheafOfModules.{v} R) noncomputable abbrev forgetToPresheafModuleCatObjObj (Y : Cᵒᵖ) : ModuleCat (R.obj X) := (ModuleCat.restrictScalars (R.map (hX.to Y))).obj (M.obj Y) -@[simp] +-- This should not be a `simp` lemma because `M.obj Y` is missing the `Module (R.obj X)` instance, +-- so `simp`ing breaks downstream proofs. lemma forgetToPresheafModuleCatObjObj_coe (Y : Cᵒᵖ) : (forgetToPresheafModuleCatObjObj X hX M Y : Type _) = M.obj Y := rfl /-- Auxiliary definition for `forgetToPresheafModuleCatObj`. -/ def forgetToPresheafModuleCatObjMap {Y Z : Cᵒᵖ} (f : Y ⟶ Z) : forgetToPresheafModuleCatObjObj X hX M Y ⟶ - forgetToPresheafModuleCatObjObj X hX M Z where - toFun x := M.map f x - map_add' := by simp - map_smul' r x := by - simp only [ModuleCat.restrictScalars.smul_def, AddHom.toFun_eq_coe, AddHom.coe_mk, - RingHom.id_apply, M.map_smul] - rw [← CategoryTheory.comp_apply, ← R.map_comp] - congr - apply hX.hom_ext + forgetToPresheafModuleCatObjObj X hX M Z := + ModuleCat.ofHom + (X := forgetToPresheafModuleCatObjObj X hX M Y) (Y := forgetToPresheafModuleCatObjObj X hX M Z) + { toFun := fun x => M.map f x + map_add' := by simp + map_smul' := fun r x => by + simp only [ModuleCat.restrictScalars.smul_def, AddHom.toFun_eq_coe, AddHom.coe_mk, + RingHom.id_apply, M.map_smul] + rw [← CategoryTheory.comp_apply, ← R.map_comp] + congr + apply hX.hom_ext } @[simp] lemma forgetToPresheafModuleCatObjMap_apply {Y Z : Cᵒᵖ} (f : Y ⟶ Z) (m : M.obj Y) : - DFunLike.coe (α := M.obj Y) (β := fun _ ↦ M.obj Z) - (forgetToPresheafModuleCatObjMap X hX M f) m = M.map f m := rfl + (forgetToPresheafModuleCatObjMap X hX M f).hom m = M.map f m := rfl /-- Implementation of the functor `PresheafOfModules R ⥤ Cᵒᵖ ⥤ ModuleCat (R.obj X)` @@ -402,10 +403,12 @@ morphism level `(f : M ⟶ N) ↦ (c ↦ f(c))`. noncomputable def forgetToPresheafModuleCatMap (X : Cᵒᵖ) (hX : Limits.IsInitial X) {M N : PresheafOfModules.{v} R} (f : M ⟶ N) : forgetToPresheafModuleCatObj X hX M ⟶ forgetToPresheafModuleCatObj X hX N where - app Y := + app Y := ModuleCat.ofHom + (X := (forgetToPresheafModuleCatObj X hX M).obj Y) + (Y := (forgetToPresheafModuleCatObj X hX N).obj Y) { toFun := f.app Y map_add' := by simp - map_smul' := fun r ↦ (f.app Y).map_smul (R.1.map (hX.to Y) _) } + map_smul' := fun r ↦ (f.app Y).hom.map_smul (R.1.map (hX.to Y) _) } naturality Y Z g := by ext x exact naturality_apply f g x diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/ChangeOfRings.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/ChangeOfRings.lean index 987578da8c709..2b6d781c0892b 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/ChangeOfRings.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/ChangeOfRings.lean @@ -28,7 +28,12 @@ variable {C : Type u'} [Category.{v'} C] {R R' : Cᵒᵖ ⥤ RingCat.{u}} noncomputable def restrictScalarsObj (M' : PresheafOfModules.{v} R') (α : R ⟶ R') : PresheafOfModules R where obj := fun X ↦ (ModuleCat.restrictScalars (α.app X)).obj (M'.obj X) - map := fun {X Y} f ↦ + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(X := ...)` and `(Y := ...)`. + -- This suggests `restrictScalars` needs to be redesigned. + map := fun {X Y} f ↦ ModuleCat.ofHom + (X := (ModuleCat.restrictScalars (α.app X)).obj (M'.obj X)) + (Y := (ModuleCat.restrictScalars (R.map f)).obj + ((ModuleCat.restrictScalars (α.app Y)).obj (M'.obj Y))) { toFun := M'.map f map_add' := map_add _ map_smul' := fun r x ↦ (M'.map_smul f (α.app _ r) x).trans (by diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean index 2b9de6dc8f4e9..f4d23fbb2a3e3 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean @@ -60,7 +60,7 @@ noncomputable def tensorObj : PresheafOfModules (R ⋙ forget₂ _ _) where intro m₁ m₂ dsimp [tensorObjMap] simp only [map_comp, Functor.comp_obj, CommRingCat.forgetToRingCat_obj, Functor.comp_map, - ModuleCat.restrictScalarsComp'_inv_app, ModuleCat.coe_comp, Function.comp_apply, + ModuleCat.restrictScalarsComp'_inv_app, ModuleCat.hom_comp, LinearMap.comp_apply, ModuleCat.restrictScalars.map_apply, ModuleCat.restrictScalarsComp'App_inv_apply] rfl) @@ -71,7 +71,7 @@ lemma tensorObj_map_tmul {X Y : Cᵒᵖ} (f : X ⟶ Y) (m₁ : M₁.obj X) (m₂ DFunLike.coe (α := (M₁.obj X ⊗ M₂.obj X : _)) (β := fun _ ↦ (ModuleCat.restrictScalars ((forget₂ CommRingCat RingCat).map (R.map f))).obj (M₁.obj Y ⊗ M₂.obj Y)) - ((tensorObj M₁ M₂).map f) (m₁ ⊗ₜ[R.obj X] m₂) = M₁.map f m₁ ⊗ₜ[R.obj Y] M₂.map f m₂ := rfl + ((tensorObj M₁ M₂).map f).hom (m₁ ⊗ₜ[R.obj X] m₂) = M₁.map f m₁ ⊗ₜ[R.obj Y] M₂.map f m₂ := rfl /-- The tensor product of two morphisms of presheaves of modules. -/ @[simps] @@ -79,10 +79,8 @@ noncomputable def tensorHom (f : M₁ ⟶ M₂) (g : M₃ ⟶ M₄) : tensorObj app X := f.app X ⊗ g.app X naturality {X Y} φ := ModuleCat.MonoidalCategory.tensor_ext (fun m₁ m₃ ↦ by dsimp - rw [tensorObj_map_tmul] - erw [ModuleCat.MonoidalCategory.tensorHom_tmul, tensorObj_map_tmul, - naturality_apply, naturality_apply] - rfl) + rw [tensorObj_map_tmul, ModuleCat.MonoidalCategory.tensorHom_tmul, tensorObj_map_tmul, + naturality_apply, naturality_apply]) end Monoidal diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Pushforward.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Pushforward.lean index ae8b9e272a34a..099ab943de7ac 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Pushforward.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Pushforward.lean @@ -37,10 +37,14 @@ def pushforward₀ (R : Dᵒᵖ ⥤ RingCat.{u}) : { obj := fun X ↦ ModuleCat.of _ (M.obj (F.op.obj X)) map := fun {X Y} f ↦ M.map (F.op.map f) map_id := fun X ↦ by - ext x + refine ModuleCat.hom_ext + -- Work around an instance diamond for `restrictScalarsId'` + (@LinearMap.ext _ _ _ _ _ _ _ _ (_) (_) _ _ _ (fun x => ?_)) exact (M.congr_map_apply (F.op.map_id X) x).trans (by simp) map_comp := fun f g ↦ by - ext x + refine ModuleCat.hom_ext + -- Work around an instance diamond for `restrictScalarsId'` + (@LinearMap.ext _ _ _ _ _ _ _ _ (_) (_) _ _ _ (fun x => ?_)) exact (M.congr_map_apply (F.op.map_comp f g) x).trans (by simp) } map {M₁ M₂} φ := { app := fun X ↦ φ.app _ } @@ -66,22 +70,29 @@ noncomputable def pushforwardCompToPresheaf : pushforward.{v} φ ⋙ toPresheaf _ ≅ toPresheaf _ ⋙ (whiskeringLeft _ _ _).obj F.op := Iso.refl _ -@[simp] lemma pushforward_obj_map_apply (M : PresheafOfModules.{v} R) {X Y : Cᵒᵖ} (f : X ⟶ Y) (m : (ModuleCat.restrictScalars (φ.app X)).obj (M.obj (Opposite.op (F.obj X.unop)))) : - DFunLike.coe - (α := (ModuleCat.restrictScalars (φ.app X)).obj (M.obj (Opposite.op (F.obj X.unop)))) - (β := fun _ ↦ (ModuleCat.restrictScalars (φ.app Y)).obj - (M.obj (Opposite.op (F.obj Y.unop)))) (((pushforward φ).obj M).map f) m = - M.map (F.map f.unop).op m := rfl + (((pushforward φ).obj M).map f).hom m = M.map (F.map f.unop).op m := rfl +/-- `@[simp]`-normal form of `pushforward_obj_map_apply`. -/ @[simp] +lemma pushforward_obj_map_apply' (M : PresheafOfModules.{v} R) {X Y : Cᵒᵖ} (f : X ⟶ Y) + (m : (ModuleCat.restrictScalars (φ.app X)).obj (M.obj (Opposite.op (F.obj X.unop)))) : + DFunLike.coe + (F := ↑((ModuleCat.restrictScalars _).obj _) →ₗ[_] + ↑((ModuleCat.restrictScalars (S.map f)).obj ((ModuleCat.restrictScalars _).obj _))) + (((pushforward φ).obj M).map f).hom m = M.map (F.map f.unop).op m := rfl + lemma pushforward_map_app_apply {M N : PresheafOfModules.{v} R} (α : M ⟶ N) (X : Cᵒᵖ) + (m : (ModuleCat.restrictScalars (φ.app X)).obj (M.obj (Opposite.op (F.obj X.unop)))) : + (((pushforward φ).map α).app X).hom m = α.app (Opposite.op (F.obj X.unop)) m := rfl + +/-- `@[simp]`-normal form of `pushforward_map_app_apply`. -/ +@[simp] +lemma pushforward_map_app_apply' {M N : PresheafOfModules.{v} R} (α : M ⟶ N) (X : Cᵒᵖ) (m : (ModuleCat.restrictScalars (φ.app X)).obj (M.obj (Opposite.op (F.obj X.unop)))) : DFunLike.coe - (α := (ModuleCat.restrictScalars (φ.app X)).obj (M.obj (Opposite.op (F.obj X.unop)))) - (β := fun _ ↦ (ModuleCat.restrictScalars (φ.app X)).obj - (N.obj (Opposite.op (F.obj X.unop)))) - (((pushforward φ).map α).app X) m = α.app (Opposite.op (F.obj X.unop)) m := rfl + (F := ↑((ModuleCat.restrictScalars _).obj _) →ₗ[_] ↑((ModuleCat.restrictScalars _).obj _)) + (((pushforward φ).map α).app X).hom m = α.app (Opposite.op (F.obj X.unop)) m := rfl end PresheafOfModules diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean index c00b766a8362d..bcf7dd4cfc6a3 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean @@ -314,10 +314,14 @@ def toSheafify : M₀ ⟶ (restrictScalars α).obj (sheafify α φ).val := simpa using (Sheafify.map_smul_eq α φ (α.app _ r₀) (φ.app _ m₀) (𝟙 _) r₀ (by aesop) m₀ (by simp)).symm) -@[simp] lemma toSheafify_app_apply (X : Cᵒᵖ) (x : M₀.obj X) : - DFunLike.coe (α := M₀.obj X) (β := fun _ ↦ A.val.obj X) - ((toSheafify α φ).app X) x = φ.app X x := rfl + ((toSheafify α φ).app X).hom x = φ.app X x := rfl + +/-- `@[simp]`-normal form of `toSheafify_app_apply`. -/ +@[simp] +lemma toSheafify_app_apply' (X : Cᵒᵖ) (x : M₀.obj X) : + DFunLike.coe (F := (_ →ₗ[_] ↑((ModuleCat.restrictScalars (α.app X)).obj _))) + ((toSheafify α φ).app X).hom x = φ.app X x := rfl @[simp] lemma toPresheaf_map_toSheafify : (toPresheaf R₀).map (toSheafify α φ) = φ := rfl @@ -370,7 +374,7 @@ def sheafifyMap (fac : (toPresheaf R₀).map τ₀ ≫ φ' = φ ≫ τ.val) : sheafify α φ ⟶ sheafify α φ' where val := homMk τ.val (fun X r m ↦ by let f := (sheafifyHomEquiv' α φ (by exact A'.cond)).symm (τ₀ ≫ toSheafify α φ') - suffices τ.val = (toPresheaf _).map f by simpa only [this] using (f.app X).map_smul r m + suffices τ.val = (toPresheaf _).map f by simpa only [this] using (f.app X).hom.map_smul r m apply ((J.W_of_isLocallyBijective φ).homEquiv _ A'.cond).injective dsimp [f] erw [comp_toPresheaf_map_sheafifyHomEquiv'_symm_hom] diff --git a/Mathlib/Algebra/Category/ModuleCat/Products.lean b/Mathlib/Algebra/Category/ModuleCat/Products.lean index 9b504d0c1dbe0..937317e6ea0a5 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Products.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Products.lean @@ -28,18 +28,16 @@ section product /-- The product cone induced by the concrete product. -/ def productCone : Fan Z := - Fan.mk (ModuleCat.of R (∀ i : ι, Z i)) fun i => (LinearMap.proj i : (∀ i : ι, Z i) →ₗ[R] Z i) + Fan.mk (ModuleCat.of R (∀ i : ι, Z i)) fun i => + ofHom (LinearMap.proj i : (∀ i : ι, Z i) →ₗ[R] Z i) /-- The concrete product cone is limiting. -/ def productConeIsLimit : IsLimit (productCone Z) where - lift s := (LinearMap.pi fun j => s.π.app ⟨j⟩ : s.pt →ₗ[R] ∀ i : ι, Z i) - fac s j := by - cases j - aesop + lift s := ofHom (LinearMap.pi fun j => (s.π.app ⟨j⟩).hom : s.pt →ₗ[R] ∀ i : ι, Z i) uniq s m w := by ext x funext i - exact LinearMap.congr_fun (w ⟨i⟩) x + exact DFunLike.congr_fun (congr_arg Hom.hom (w ⟨i⟩)) x -- While we could use this to construct a `HasProducts (ModuleCat R)` instance, -- we already have `HasLimits (ModuleCat R)` in `Algebra.Category.ModuleCat.Limits`. @@ -54,12 +52,12 @@ noncomputable def piIsoPi : ∏ᶜ Z ≅ ModuleCat.of R (∀ i, Z i) := -- We now show this isomorphism commutes with the inclusion of the kernel into the source. @[simp, elementwise] theorem piIsoPi_inv_kernel_ι (i : ι) : - (piIsoPi Z).inv ≫ Pi.π Z i = (LinearMap.proj i : (∀ i : ι, Z i) →ₗ[R] Z i) := + (piIsoPi Z).inv ≫ Pi.π Z i = ofHom (LinearMap.proj i : (∀ i : ι, Z i) →ₗ[R] Z i) := limit.isoLimitCone_inv_π _ _ @[simp, elementwise] theorem piIsoPi_hom_ker_subtype (i : ι) : - (piIsoPi Z).hom ≫ (LinearMap.proj i : (∀ i : ι, Z i) →ₗ[R] Z i) = Pi.π Z i := + (piIsoPi Z).hom ≫ ofHom (LinearMap.proj i : (∀ i : ι, Z i) →ₗ[R] Z i) = Pi.π Z i := IsLimit.conePointUniqueUpToIso_inv_comp _ (limit.isLimit _) (Discrete.mk i) end product @@ -72,19 +70,20 @@ variable [DecidableEq ι] /-- The coproduct cone induced by the concrete product. -/ def coproductCocone : Cofan Z := - Cofan.mk (ModuleCat.of R (⨁ i : ι, Z i)) fun i => (DirectSum.lof R ι (fun i ↦ Z i) i) + Cofan.mk (ModuleCat.of R (⨁ i : ι, Z i)) fun i => ofHom (DirectSum.lof R ι (fun i ↦ Z i) i) /-- The concrete coproduct cone is limiting. -/ def coproductCoconeIsColimit : IsColimit (coproductCocone Z) where - desc s := DirectSum.toModule R ι _ fun i ↦ s.ι.app ⟨i⟩ + desc s := ofHom <| DirectSum.toModule R ι _ fun i ↦ (s.ι.app ⟨i⟩).hom fac := by rintro s ⟨i⟩ ext (x : Z i) simpa only [Discrete.functor_obj_eq_as, coproductCocone, Cofan.mk_pt, Functor.const_obj_obj, - Cofan.mk_ι_app, coe_comp, Function.comp_apply] using + Cofan.mk_ι_app, hom_comp, LinearMap.coe_comp, Function.comp_apply] using DirectSum.toModule_lof (ι := ι) R (M := fun i ↦ Z i) i x uniq := by rintro s f h + ext : 1 refine DirectSum.linearMap_ext _ fun i ↦ ?_ ext x simpa only [LinearMap.coe_comp, Function.comp_apply, toModule_lof] using @@ -100,12 +99,12 @@ noncomputable def coprodIsoDirectSum : ∐ Z ≅ ModuleCat.of R (⨁ i, Z i) := @[simp, elementwise] theorem ι_coprodIsoDirectSum_hom (i : ι) : - Sigma.ι Z i ≫ (coprodIsoDirectSum Z).hom = DirectSum.lof R ι (fun i ↦ Z i) i := + Sigma.ι Z i ≫ (coprodIsoDirectSum Z).hom = ofHom (DirectSum.lof R ι (fun i ↦ Z i) i) := colimit.isoColimitCocone_ι_hom _ _ @[simp, elementwise] theorem lof_coprodIsoDirectSum_inv (i : ι) : - DirectSum.lof R ι (fun i ↦ Z i) i ≫ (coprodIsoDirectSum Z).inv = Sigma.ι Z i := + ofHom (DirectSum.lof R ι (fun i ↦ Z i) i) ≫ (coprodIsoDirectSum Z).inv = Sigma.ι Z i := (coproductCoconeIsColimit Z).comp_coconePointUniqueUpToIso_hom (colimit.isColimit _) _ end coproduct diff --git a/Mathlib/Algebra/Category/ModuleCat/Projective.lean b/Mathlib/Algebra/Category/ModuleCat/Projective.lean index e8bfc43488358..1947ab1f393c3 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Projective.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Projective.lean @@ -30,13 +30,16 @@ theorem IsProjective.iff_projective {R : Type u} [Ring R] {P : Type max u v} [Ad [Module R P] : Module.Projective R P ↔ Projective (ModuleCat.of R P) := by refine ⟨fun h => ?_, fun h => ?_⟩ · letI : Module.Projective R (ModuleCat.of R P) := h - exact ⟨fun E X epi => Module.projective_lifting_property _ _ - ((ModuleCat.epi_iff_surjective _).mp epi)⟩ + refine ⟨fun E X epi => ?_⟩ + obtain ⟨f, h⟩ := Module.projective_lifting_property X.hom E.hom + ((ModuleCat.epi_iff_surjective _).mp epi) + exact ⟨ofHom f, hom_ext h⟩ · refine Module.Projective.of_lifting_property.{u,v} ?_ intro E X mE mX sE sX f g s haveI : Epi (↟f) := (ModuleCat.epi_iff_surjective (↟f)).mpr s letI : Projective (ModuleCat.of R P) := h - exact ⟨Projective.factorThru (↟g) (↟f), Projective.factorThru_comp (↟g) (↟f)⟩ + exact ⟨(Projective.factorThru (↟g) (↟f)).hom, + ModuleCat.hom_ext_iff.mp <| Projective.factorThru_comp (↟g) (↟f)⟩ namespace ModuleCat @@ -56,15 +59,9 @@ instance moduleCat_enoughProjectives : EnoughProjectives (ModuleCat.{max u v} R) projective := projective_of_free.{v,u} (ι := M) (M := ModuleCat.of R (M →₀ R)) <| Finsupp.basisSingleOne - f := Finsupp.basisSingleOne.constr ℕ _root_.id + f := ofHom <| Finsupp.basisSingleOne.constr ℕ _root_.id epi := (epi_iff_range_eq_top _).mpr (range_eq_top.2 fun m => ⟨Finsupp.single m (1 : R), by - -- Porting note: simp [Finsupp.linearCombination_single] fails but rw succeeds - dsimp [Basis.constr] - simp only [Finsupp.lmapDomain_id, comp_id] - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [Finsupp.linearCombination_single] - rw [one_smul] - rfl ⟩) }⟩ + simp [Finsupp.linearCombination_single, Basis.constr] ⟩)}⟩ end ModuleCat diff --git a/Mathlib/Algebra/Category/ModuleCat/Sheaf.lean b/Mathlib/Algebra/Category/ModuleCat/Sheaf.lean index 3a23ba4b4c02e..cf467ff8b32d1 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Sheaf.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Sheaf.lean @@ -219,7 +219,7 @@ noncomputable def homEquivOfIsLocallyBijective : (M₂ ⟶ N) ≃ (M₁ ⟶ N) w have hφ' : ∀ (z : M₂.obj X), φ.app _ (M₂.map p.op z) = N.map p.op (φ.app _ z) := congr_fun ((forget _).congr_map (φ.naturality p.op)) change N.map p.op (φ.app X (r • y)) = N.map p.op (r • φ.app X y) - rw [← hφ', M₂.map_smul, ← hx, ← (f.app _).map_smul, hφ, (ψ.app _).map_smul, + rw [← hφ', M₂.map_smul, ← hx, ← (f.app _).hom.map_smul, hφ, (ψ.app _).hom.map_smul, ← hφ, hx, N.map_smul, hφ']) left_inv φ := (toPresheaf _).map_injective (((J.W_of_isLocallyBijective diff --git a/Mathlib/Algebra/Category/ModuleCat/Sheaf/ChangeOfRings.lean b/Mathlib/Algebra/Category/ModuleCat/Sheaf/ChangeOfRings.lean index 53503eb352661..a73bfc742c383 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Sheaf/ChangeOfRings.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Sheaf/ChangeOfRings.lean @@ -59,7 +59,7 @@ noncomputable def restrictHomEquivOfIsLocallySurjective change M₂.map p.op (g.app X (r' • m)) = M₂.map p.op (r' • show M₂.obj X from g.app X m) dsimp at hg ⊢ rw [← hg, M₂.map_smul, ← hg, ← hr] - erw [← (g.app _).map_smul] + erw [← (g.app _).hom.map_smul] rw [M₁.map_smul, ← hr] rfl) left_inv _ := rfl diff --git a/Mathlib/Algebra/Category/ModuleCat/Subobject.lean b/Mathlib/Algebra/Category/ModuleCat/Subobject.lean index 0facef08450f1..3c70d7b6228ef 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Subobject.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Subobject.lean @@ -35,39 +35,37 @@ variable {R : Type u} [Ring R] (M : ModuleCat.{v} R) submodules. -/ noncomputable def subobjectModule : Subobject M ≃o Submodule R M := OrderIso.symm - { invFun := fun S => LinearMap.range S.arrow - toFun := fun N => Subobject.mk (↾N.subtype) + { invFun := fun S => LinearMap.range S.arrow.hom + toFun := fun N => Subobject.mk (ofHom N.subtype) right_inv := fun S => Eq.symm (by fapply eq_mk_of_comm - · apply LinearEquiv.toModuleIso'Left - apply LinearEquiv.ofBijective (LinearMap.codRestrict (LinearMap.range S.arrow) S.arrow _) + · apply LinearEquiv.toModuleIso + apply LinearEquiv.ofBijective (LinearMap.codRestrict + (LinearMap.range S.arrow.hom) S.arrow.hom _) constructor · simp [← LinearMap.ker_eq_bot, LinearMap.ker_codRestrict] rw [ker_eq_bot_of_mono] · rw [← LinearMap.range_eq_top, LinearMap.range_codRestrict, Submodule.comap_subtype_self] exact LinearMap.mem_range_self _ - · apply LinearMap.ext - intro x + · ext x rfl) left_inv := fun N => by - -- Porting note: The type of `↾N.subtype` was ambiguous. Not entirely sure, I made the right - -- choice here - convert congr_arg LinearMap.range - (underlyingIso_arrow (↾N.subtype : of R { x // x ∈ N } ⟶ M)) using 1 + convert congr_arg LinearMap.range (ModuleCat.hom_ext_iff.mp + (underlyingIso_arrow (ofHom N.subtype))) using 1 · have : -- Porting note: added the `.toLinearEquiv.toLinearMap` - (underlyingIso (↾N.subtype : of R _ ⟶ M)).inv = - (underlyingIso (↾N.subtype : of R _ ⟶ M)).symm.toLinearEquiv.toLinearMap := by - apply LinearMap.ext - intro x + (underlyingIso (ofHom N.subtype)).inv = + ofHom (underlyingIso (ofHom N.subtype)).symm.toLinearEquiv.toLinearMap := by + ext x rfl - rw [this, comp_def, LinearEquiv.range_comp] + rw [this, hom_comp, LinearEquiv.range_comp] · exact (Submodule.range_subtype _).symm map_rel_iff' := fun {S T} => by refine ⟨fun h => ?_, fun h => mk_le_mk_of_comm (↟(Submodule.inclusion h)) rfl⟩ - convert LinearMap.range_comp_le_range (ofMkLEMk _ _ h) (↾T.subtype) - · simpa only [← comp_def, ofMkLEMk_comp] using (Submodule.range_subtype _).symm + convert LinearMap.range_comp_le_range (ofMkLEMk _ _ h).hom (ofHom T.subtype).hom + · rw [← hom_comp, ofMkLEMk_comp] + exact (Submodule.range_subtype _).symm · exact (Submodule.range_subtype _).symm } instance wellPowered_moduleCat : WellPowered (ModuleCat.{v} R) := @@ -77,16 +75,16 @@ attribute [local instance] hasKernels_moduleCat /-- Bundle an element `m : M` such that `f m = 0` as a term of `kernelSubobject f`. -/ noncomputable def toKernelSubobject {M N : ModuleCat.{v} R} {f : M ⟶ N} : - LinearMap.ker f →ₗ[R] kernelSubobject f := - (kernelSubobjectIso f ≪≫ ModuleCat.kernelIsoKer f).inv + LinearMap.ker f.hom →ₗ[R] kernelSubobject f := + (kernelSubobjectIso f ≪≫ ModuleCat.kernelIsoKer f).inv.hom @[simp] -theorem toKernelSubobject_arrow {M N : ModuleCat R} {f : M ⟶ N} (x : LinearMap.ker f) : +theorem toKernelSubobject_arrow {M N : ModuleCat R} {f : M ⟶ N} (x : LinearMap.ker f.hom) : (kernelSubobject f).arrow (toKernelSubobject x) = x.1 := by -- Porting note (https://github.com/leanprover-community/mathlib4/pull/10959): the whole proof was just `simp [toKernelSubobject]`. suffices ((arrow ((kernelSubobject f))) ∘ (kernelSubobjectIso f ≪≫ kernelIsoKer f).inv) x = x by convert this - rw [Iso.trans_inv, ← coe_comp, Category.assoc] + rw [Iso.trans_inv, ← LinearMap.coe_comp, ← hom_comp, Category.assoc] simp only [Category.assoc, kernelSubobject_arrow', kernelIsoKer_inv_kernel_ι] aesop_cat @@ -108,9 +106,13 @@ theorem cokernel_π_imageSubobject_ext {L M N : ModuleCat.{v} R} (f : L ⟶ M) [ subst w -- Porting note (https://github.com/leanprover-community/mathlib4/pull/10959): The proof from here used to just be `simp`. simp only [map_add, add_right_eq_self] - change ((cokernel.π g) ∘ (g) ∘ (factorThruImageSubobject f)) l = 0 - rw [← coe_comp, ← coe_comp, Category.assoc] - simp only [cokernel.condition, comp_zero] - rfl + -- TODO: add a `@[simp]` lemma along the lines of: + -- ``` + -- lemma ModuleCat.Hom.cokernel_condition : (cokernel.π g).hom (g.hom x) = 0 + -- ``` + -- ideally generated for all concrete categories (using a metaprogram like `@[elementwise]`?). + -- See also: https://github.com/leanprover-community/mathlib4/pull/19511#discussion_r1867083077 + change ((factorThruImageSubobject f) ≫ g ≫ (cokernel.π g)).hom l = 0 + simp end ModuleCat diff --git a/Mathlib/Algebra/Category/ModuleCat/Tannaka.lean b/Mathlib/Algebra/Category/ModuleCat/Tannaka.lean index a22898dd2560d..febe6bfa7f0fc 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Tannaka.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Tannaka.lean @@ -26,10 +26,7 @@ def ringEquivEndForget₂ (R : Type u) [Ring R] : R ≃+* End (AdditiveFunctor.of (forget₂ (ModuleCat.{u} R) AddCommGrp.{u})) where toFun r := { app := fun M => - @AddCommGrp.ofHom M.carrier M.carrier _ _ (DistribMulAction.toAddMonoidHom M r) - naturality := fun M N f => by - ext - exact (f.map_smul _ _).symm } + @AddCommGrp.ofHom M.carrier M.carrier _ _ (DistribMulAction.toAddMonoidHom M r) } invFun φ := φ.app (ModuleCat.of R R) (1 : R) left_inv := by intro r @@ -39,19 +36,19 @@ def ringEquivEndForget₂ (R : Type u) [Ring R] : apply NatTrans.ext ext M (x : M) have w := congr_fun ((forget _).congr_map - (φ.naturality (ModuleCat.asHomRight (LinearMap.toSpanSingleton R M x)))) (1 : R) + (φ.naturality (ModuleCat.ofHom (LinearMap.toSpanSingleton R M x)))) (1 : R) exact w.symm.trans (congr_arg (φ.app M) (one_smul R x)) map_add' := by intros apply NatTrans.ext ext - dsimp - simp only [AddCommGrp.ofHom_apply, DistribMulAction.toAddMonoidHom_apply, add_smul] + simp only [AdditiveFunctor.of_fst, ModuleCat.forget₂_obj, AddCommGrp.coe_of, + AddCommGrp.ofHom_apply, DistribMulAction.toAddMonoidHom_apply, add_smul] rfl map_mul' := by intros apply NatTrans.ext ext - dsimp - simp only [AddCommGrp.ofHom_apply, DistribMulAction.toAddMonoidHom_apply, mul_smul] + simp only [AdditiveFunctor.of_fst, ModuleCat.forget₂_obj, AddCommGrp.coe_of, + AddCommGrp.ofHom_apply, DistribMulAction.toAddMonoidHom_apply, mul_smul] rfl diff --git a/Mathlib/Algebra/Homology/LocalCohomology.lean b/Mathlib/Algebra/Homology/LocalCohomology.lean index a66ee204117c8..c98cac0916bc0 100644 --- a/Mathlib/Algebra/Homology/LocalCohomology.lean +++ b/Mathlib/Algebra/Homology/LocalCohomology.lean @@ -66,9 +66,7 @@ variable {R : Type u} [CommRing R] {D : Type v} [SmallCategory D] determined by the functor `I` -/ def ringModIdeals (I : D ⥤ Ideal R) : D ⥤ ModuleCat.{u} R where obj t := ModuleCat.of R <| R ⧸ I.obj t - map w := Submodule.mapQ _ _ LinearMap.id (I.map w).down.down - -- Porting note: was 'obviously' - map_comp f g := by apply Submodule.linearMap_qext; rfl + map w := ModuleCat.ofHom <| Submodule.mapQ _ _ LinearMap.id (I.map w).down.down -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11215): TODO: Once this file is ported, move this instance to the right location. instance moduleCat_enoughProjectives' : EnoughProjectives (ModuleCat.{u} R) := diff --git a/Mathlib/Algebra/Homology/ShortComplex/ModuleCat.lean b/Mathlib/Algebra/Homology/ShortComplex/ModuleCat.lean index 7523eadd214dd..87236bc52bb4d 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/ModuleCat.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/ModuleCat.lean @@ -33,7 +33,7 @@ linear maps `f` and `g` and the vanishing of their composition. -/ def moduleCatMk {X₁ X₂ X₃ : Type v} [AddCommGroup X₁] [AddCommGroup X₂] [AddCommGroup X₃] [Module R X₁] [Module R X₂] [Module R X₃] (f : X₁ →ₗ[R] X₂) (g : X₂ →ₗ[R] X₃) (hfg : g.comp f = 0) : ShortComplex (ModuleCat.{v} R) := - ShortComplex.mk (ModuleCat.ofHom f) (ModuleCat.ofHom g) hfg + ShortComplex.mk (ModuleCat.ofHom f) (ModuleCat.ofHom g) (ModuleCat.hom_ext hfg) variable (S : ShortComplex (ModuleCat.{v} R)) @@ -46,33 +46,19 @@ lemma moduleCat_exact_iff : S.exact_iff_of_concreteCategory lemma moduleCat_exact_iff_ker_sub_range : - S.Exact ↔ LinearMap.ker S.g ≤ LinearMap.range S.f := by + S.Exact ↔ LinearMap.ker S.g.hom ≤ LinearMap.range S.f.hom := by rw [moduleCat_exact_iff] - constructor - · intro h x₂ hx₂ - exact h x₂ hx₂ - · intro h x₂ hx₂ - exact h hx₂ + aesop lemma moduleCat_exact_iff_range_eq_ker : - S.Exact ↔ LinearMap.range S.f = LinearMap.ker S.g := by + S.Exact ↔ LinearMap.range S.f.hom = LinearMap.ker S.g.hom := by rw [moduleCat_exact_iff_ker_sub_range] - constructor - · intro h - ext x - constructor - · rintro ⟨y, hy⟩ - rw [← hy] - simp only [LinearMap.mem_ker, moduleCat_zero_apply] - · intro hx - exact h hx - · intro h - rw [h] + aesop variable {S} lemma Exact.moduleCat_range_eq_ker (hS : S.Exact) : - LinearMap.range S.f = LinearMap.ker S.g := by + LinearMap.range S.f.hom = LinearMap.ker S.g.hom := by simpa only [moduleCat_exact_iff_range_eq_ker] using hS lemma ShortExact.moduleCat_injective_f (hS : S.ShortExact) : @@ -94,19 +80,17 @@ lemma ShortExact.moduleCat_exact_iff_function_exact : morphisms `f` and `g` and the assumption `LinearMap.range f ≤ LinearMap.ker g`. -/ @[simps] def moduleCatMkOfKerLERange {X₁ X₂ X₃ : ModuleCat.{v} R} (f : X₁ ⟶ X₂) (g : X₂ ⟶ X₃) - (hfg : LinearMap.range f ≤ LinearMap.ker g) : ShortComplex (ModuleCat.{v} R) := - ShortComplex.mk f g (by - ext - exact hfg ⟨_, rfl⟩) + (hfg : LinearMap.range f.hom ≤ LinearMap.ker g.hom) : ShortComplex (ModuleCat.{v} R) := + ShortComplex.mk f g (by aesop) lemma Exact.moduleCat_of_range_eq_ker {X₁ X₂ X₃ : ModuleCat.{v} R} - (f : X₁ ⟶ X₂) (g : X₂ ⟶ X₃) (hfg : LinearMap.range f = LinearMap.ker g) : + (f : X₁ ⟶ X₂) (g : X₂ ⟶ X₃) (hfg : LinearMap.range f.hom = LinearMap.ker g.hom) : (moduleCatMkOfKerLERange f g (by rw [hfg])).Exact := by simpa only [moduleCat_exact_iff_range_eq_ker] using hfg /-- The canonical linear map `S.X₁ →ₗ[R] LinearMap.ker S.g` induced by `S.f`. -/ @[simps] -def moduleCatToCycles : S.X₁ →ₗ[R] LinearMap.ker S.g where +def moduleCatToCycles : S.X₁ →ₗ[R] LinearMap.ker S.g.hom where toFun x := ⟨S.f x, S.moduleCat_zero_apply x⟩ map_add' x y := by aesop map_smul' a x := by aesop @@ -114,35 +98,28 @@ def moduleCatToCycles : S.X₁ →ₗ[R] LinearMap.ker S.g where /-- The homology of `S`, defined as the quotient of the kernel of `S.g` by the image of `S.moduleCatToCycles` -/ abbrev moduleCatHomology := - ModuleCat.of R (LinearMap.ker S.g ⧸ LinearMap.range S.moduleCatToCycles) + ModuleCat.of R (LinearMap.ker S.g.hom ⧸ LinearMap.range S.moduleCatToCycles) /-- The canonical map `ModuleCat.of R (LinearMap.ker S.g) ⟶ S.moduleCatHomology`. -/ -abbrev moduleCatHomologyπ : ModuleCat.of R (LinearMap.ker S.g) ⟶ S.moduleCatHomology := - (LinearMap.range S.moduleCatToCycles).mkQ +abbrev moduleCatHomologyπ : ModuleCat.of R (LinearMap.ker S.g.hom) ⟶ S.moduleCatHomology := + ModuleCat.ofHom (LinearMap.range S.moduleCatToCycles).mkQ /-- The explicit left homology data of a short complex of modules that is given by a kernel and a quotient given by the `LinearMap` API. -/ -@[simps] +@[simps K H i π] def moduleCatLeftHomologyData : S.LeftHomologyData where - K := ModuleCat.of R (LinearMap.ker S.g) + K := ModuleCat.of R (LinearMap.ker S.g.hom) H := S.moduleCatHomology - i := (LinearMap.ker S.g).subtype + i := ModuleCat.ofHom (LinearMap.ker S.g.hom).subtype π := S.moduleCatHomologyπ - wi := by - ext ⟨_, hx⟩ - exact hx + wi := by aesop hi := ModuleCat.kernelIsLimit _ - wπ := by - ext (x : S.X₁) - dsimp - erw [Submodule.Quotient.mk_eq_zero] - rw [LinearMap.mem_range] - apply exists_apply_eq_apply + wπ := by aesop hπ := ModuleCat.cokernelIsColimit (ModuleCat.ofHom S.moduleCatToCycles) @[simp] lemma moduleCatLeftHomologyData_f' : - S.moduleCatLeftHomologyData.f' = S.moduleCatToCycles := rfl + S.moduleCatLeftHomologyData.f' = ModuleCat.ofHom S.moduleCatToCycles := rfl instance : Epi S.moduleCatHomologyπ := (inferInstance : Epi S.moduleCatLeftHomologyData.π) @@ -150,22 +127,22 @@ instance : Epi S.moduleCatHomologyπ := /-- Given a short complex `S` of modules, this is the isomorphism between the abstract `S.cycles` of the homology API and the more concrete description as `LinearMap.ker S.g`. -/ -noncomputable def moduleCatCyclesIso : S.cycles ≅ ModuleCat.of R (LinearMap.ker S.g) := +noncomputable def moduleCatCyclesIso : S.cycles ≅ ModuleCat.of R (LinearMap.ker S.g.hom) := S.moduleCatLeftHomologyData.cyclesIso @[reassoc (attr := simp, elementwise)] lemma moduleCatCyclesIso_hom_subtype : - S.moduleCatCyclesIso.hom ≫ (LinearMap.ker S.g).subtype = S.iCycles := + S.moduleCatCyclesIso.hom ≫ ModuleCat.ofHom (LinearMap.ker S.g.hom).subtype = S.iCycles := S.moduleCatLeftHomologyData.cyclesIso_hom_comp_i @[reassoc (attr := simp, elementwise)] lemma moduleCatCyclesIso_inv_iCycles : - S.moduleCatCyclesIso.inv ≫ S.iCycles = (LinearMap.ker S.g).subtype := + S.moduleCatCyclesIso.inv ≫ S.iCycles = ModuleCat.ofHom (LinearMap.ker S.g.hom).subtype := S.moduleCatLeftHomologyData.cyclesIso_inv_comp_iCycles @[reassoc (attr := simp, elementwise)] lemma toCycles_moduleCatCyclesIso_hom : - S.toCycles ≫ S.moduleCatCyclesIso.hom = S.moduleCatToCycles := by + S.toCycles ≫ S.moduleCatCyclesIso.hom = ModuleCat.ofHom S.moduleCatToCycles := by rw [← cancel_mono S.moduleCatLeftHomologyData.i, moduleCatLeftHomologyData_i, Category.assoc, S.moduleCatCyclesIso_hom_subtype, toCycles_i] rfl diff --git a/Mathlib/AlgebraicGeometry/AffineSpace.lean b/Mathlib/AlgebraicGeometry/AffineSpace.lean index 7751d57d64eda..ac5248a0b208b 100644 --- a/Mathlib/AlgebraicGeometry/AffineSpace.lean +++ b/Mathlib/AlgebraicGeometry/AffineSpace.lean @@ -295,9 +295,7 @@ lemma map_comp {S S' S'' : Scheme} (f : S ⟶ S') (g : S' ⟶ S'') : map n (f ≫ g) = map n f ≫ map n g := by ext1 · simp - · simp only [TopologicalSpace.Opens.map_top, Scheme.comp_coeBase, - TopologicalSpace.Opens.map_comp_obj, Scheme.comp_app, CommRingCat.comp_apply] - rw [map_appTop_coord, map_appTop_coord, map_appTop_coord] + · simp lemma map_Spec_map {R S : CommRingCat.{max u v}} (φ : R ⟶ S) : map n (Spec.map φ) = diff --git a/Mathlib/AlgebraicGeometry/Modules/Tilde.lean b/Mathlib/AlgebraicGeometry/Modules/Tilde.lean index 12264f5de8c43..763031d694da3 100644 --- a/Mathlib/AlgebraicGeometry/Modules/Tilde.lean +++ b/Mathlib/AlgebraicGeometry/Modules/Tilde.lean @@ -158,7 +158,11 @@ noncomputable instance (U : (Opens (PrimeSpectrum.Top R))ᵒᵖ) : noncomputable def tilde : (Spec (CommRingCat.of R)).Modules where val := { obj := fun U ↦ ModuleCat.of _ (M.tildeInType.val.obj U) - map := fun {U V} i ↦ + map := fun {U V} i ↦ ofHom + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(Y := ...)` + -- This suggests `restrictScalars` needs to be redesigned. + (Y := (restrictScalars ((Spec (CommRingCat.of R)).ringCatSheaf.val.map i)).obj + (of ((Spec (CommRingCat.of R)).ringCatSheaf.val.obj V) (M.tildeInType.val.obj V))) { toFun := M.tildeInType.val.map i map_smul' := by intros; rfl map_add' := by intros; rfl } } @@ -200,15 +204,18 @@ If `U` is an open subset of `Spec R`, this is the morphism of `R`-modules from ` `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 + ModuleCat.of R M ⟶ (tildeInModuleCat M).1.obj (op U) := + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(Y := ...)` + -- This suggests `restrictScalars` needs to be redesigned. + ModuleCat.ofHom (Y := (tildeInModuleCat M).1.obj (op U)) + { toFun := fun 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' := fun f g => Subtype.eq <| funext fun x ↦ LinearMap.map_add _ _ _ + map_smul' := fun 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) : @@ -235,7 +242,8 @@ lemma isUnit_toStalk (x : PrimeSpectrum.Top R) (r : x.asIdeal.primeCompl) : ⟨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) _ <| + germ_ext (C := ModuleCat R) (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 ↦ ?_⟩ @@ -252,7 +260,7 @@ 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 + ModuleCat.ofHom <| LocalizedModule.lift _ (toStalk M x).hom <| isUnit_toStalk M x /-- The ring homomorphism that takes a section of the structure sheaf of `R` on the open set `U`, @@ -260,10 +268,15 @@ implemented as a subtype of dependent functions to localizations at prime ideals the section on the point corresponding to a given prime ideal. -/ def openToLocalization (U : Opens (PrimeSpectrum R)) (x : PrimeSpectrum R) (hx : x ∈ U) : (tildeInModuleCat M).obj (op U) ⟶ - ModuleCat.of R (LocalizedModule x.asIdeal.primeCompl M) where - toFun s := (s.1 ⟨x, hx⟩ : _) - map_add' _ _ := rfl - map_smul' _ _ := rfl + ModuleCat.of R (LocalizedModule x.asIdeal.primeCompl M) := + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(X := ...)` and `(Y := ...)` + -- This suggests `restrictScalars` needs to be redesigned. + ModuleCat.ofHom + (X := (tildeInModuleCat M).obj (op U)) + (Y := ModuleCat.of R (LocalizedModule x.asIdeal.primeCompl M)) + { toFun := fun s => (s.1 ⟨x, hx⟩ : _) + map_add' := fun _ _ => rfl + map_smul' := fun _ _ => rfl } /-- The morphism of `R`-modules from the stalk of `M^~` at `x` to the localization of `M` at the @@ -286,9 +299,9 @@ theorem germ_comp_stalkToFiberLinearMap (U : Opens (PrimeSpectrum.Top R)) (x) (h @[simp] theorem stalkToFiberLinearMap_germ (U : Opens (PrimeSpectrum.Top R)) (x : PrimeSpectrum.Top R) (hx : x ∈ U) (s : (tildeInModuleCat M).1.obj (op U)) : - stalkToFiberLinearMap M x + (stalkToFiberLinearMap M x).hom (TopCat.Presheaf.germ (tildeInModuleCat M) U x hx s) = (s.1 ⟨x, hx⟩ : _) := - DFunLike.ext_iff.1 (germ_comp_stalkToFiberLinearMap M U x hx) s + DFunLike.ext_iff.1 (ModuleCat.hom_ext_iff.mp (germ_comp_stalkToFiberLinearMap M U x hx)) s @[reassoc (attr := simp), elementwise (attr := simp)] theorem toOpen_germ (U : Opens (PrimeSpectrum.Top R)) (x) (hx : x ∈ U) : @@ -298,13 +311,13 @@ theorem toOpen_germ (U : Opens (PrimeSpectrum.Top R)) (x) (hx : x ∈ U) : @[reassoc (attr := simp)] theorem toStalk_comp_stalkToFiberLinearMap (x : PrimeSpectrum.Top R) : toStalk M x ≫ stalkToFiberLinearMap M x = - LocalizedModule.mkLinearMap x.asIdeal.primeCompl M := by + ofHom (LocalizedModule.mkLinearMap x.asIdeal.primeCompl M) := by rw [toStalk, Category.assoc, germ_comp_stalkToFiberLinearMap]; rfl theorem stalkToFiberLinearMap_toStalk (x : PrimeSpectrum.Top R) (m : M) : - (stalkToFiberLinearMap M x) (toStalk M x m) = + (stalkToFiberLinearMap M x).hom (toStalk M x m) = LocalizedModule.mk m 1 := - LinearMap.ext_iff.1 (toStalk_comp_stalkToFiberLinearMap M x) _ + LinearMap.ext_iff.1 (ModuleCat.hom_ext_iff.mp (toStalk_comp_stalkToFiberLinearMap M x)) _ /-- If `U` is an open subset of `Spec R`, `m` is an element of `M` and `r` is an element of `R` @@ -347,7 +360,7 @@ theorem res_const (f : M) (g : R) (U hu V hv i) : @[simp] theorem localizationToStalk_mk (x : PrimeSpectrum.Top R) (f : M) (s : x.asIdeal.primeCompl) : - localizationToStalk M x (LocalizedModule.mk f s) = + (localizationToStalk M x).hom (LocalizedModule.mk f s) = (tildeInModuleCat M).germ (PrimeSpectrum.basicOpen (s : R)) x s.2 (const M f s (PrimeSpectrum.basicOpen s) fun _ => id) := (Module.End_isUnit_iff _ |>.1 (isUnit_toStalk M x s)).injective <| by @@ -356,6 +369,7 @@ theorem localizationToStalk_mk (x : PrimeSpectrum.Top R) (f : M) (s : x.asIdeal. show (M.tildeInModuleCat.germ ⊤ x ⟨⟩) ((toOpen M ⊤) f) = _ rw [← map_smul] fapply TopCat.Presheaf.germ_ext (W := PrimeSpectrum.basicOpen s.1) (hxW := s.2) + (F := M.tildeInModuleCat) · exact homOfLE le_top · exact 𝟙 _ refine Subtype.eq <| funext fun y => show LocalizedModule.mk f 1 = _ from ?_ @@ -376,7 +390,8 @@ noncomputable def stalkIso (x : PrimeSpectrum.Top R) : ModuleCat.of R (LocalizedModule x.asIdeal.primeCompl M) where hom := stalkToFiberLinearMap M x inv := localizationToStalk M x - hom_inv_id := TopCat.Presheaf.stalk_hom_ext _ fun U hxU ↦ ext _ fun s ↦ by + hom_inv_id := TopCat.Presheaf.stalk_hom_ext _ fun U hxU ↦ ModuleCat.hom_ext <| + LinearMap.ext fun s ↦ by show localizationToStalk M x (stalkToFiberLinearMap M x (M.tildeInModuleCat.germ U x hxU s)) = M.tildeInModuleCat.germ U x hxU s rw [stalkToFiberLinearMap_germ] @@ -384,15 +399,19 @@ noncomputable def stalkIso (x : PrimeSpectrum.Top R) : exists_const _ _ s x hxU rw [← res_apply M U V iVU s ⟨x, hxV⟩, ← hs, const_apply, localizationToStalk_mk] exact (tildeInModuleCat M).germ_ext V hxV (homOfLE hg) iVU <| hs ▸ rfl - inv_hom_id := by ext x; exact x.induction_on (fun _ _ => by simp) + inv_hom_id := by ext x; exact x.induction_on (fun _ _ => by + simp only [hom_comp, LinearMap.coe_comp, Function.comp_apply, hom_id, LinearMap.id_coe, id_eq] + rw [localizationToStalk_mk, stalkToFiberLinearMap_germ] + simp) instance (x : PrimeSpectrum.Top R) : - IsLocalizedModule x.asIdeal.primeCompl (toStalk M x) := by + IsLocalizedModule x.asIdeal.primeCompl (toStalk M x).hom := by convert IsLocalizedModule.of_linearEquiv (hf := localizedModuleIsLocalizedModule (M := M) x.asIdeal.primeCompl) (e := (stalkIso M x).symm.toLinearEquiv) - simp only [of_coe, show (stalkIso M x).symm.toLinearEquiv.toLinearMap = (stalkIso M x).inv by rfl, - stalkIso_inv] + ext + simp only [of_coe, + show (stalkIso M x).symm.toLinearEquiv.toLinearMap = (stalkIso M x).inv.hom by rfl] erw [LocalizedModule.lift_comp] end Tilde diff --git a/Mathlib/CategoryTheory/Abelian/Pseudoelements.lean b/Mathlib/CategoryTheory/Abelian/Pseudoelements.lean index 447bcadee24a4..8aa0fd2f01e4b 100644 --- a/Mathlib/CategoryTheory/Abelian/Pseudoelements.lean +++ b/Mathlib/CategoryTheory/Abelian/Pseudoelements.lean @@ -427,20 +427,20 @@ section Module /-- In the category `Module R`, if `x` and `y` are pseudoequal, then the range of the associated morphisms is the same. -/ theorem ModuleCat.eq_range_of_pseudoequal {R : Type*} [CommRing R] {G : ModuleCat R} {x y : Over G} - (h : PseudoEqual G x y) : LinearMap.range x.hom = LinearMap.range y.hom := by + (h : PseudoEqual G x y) : LinearMap.range x.hom.hom = LinearMap.range y.hom.hom := by obtain ⟨P, p, q, hp, hq, H⟩ := h refine Submodule.ext fun a => ⟨fun ha => ?_, fun ha => ?_⟩ · obtain ⟨a', ha'⟩ := ha obtain ⟨a'', ha''⟩ := (ModuleCat.epi_iff_surjective p).1 hp a' refine ⟨q a'', ?_⟩ - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [← LinearMap.comp_apply, ← ModuleCat.comp_def, ← H, - ModuleCat.comp_def, LinearMap.comp_apply, ha'', ha'] + dsimp at ha' ⊢ + rw [← LinearMap.comp_apply, ← ModuleCat.hom_comp, ← H, + ModuleCat.hom_comp, LinearMap.comp_apply, ha'', ha'] · obtain ⟨a', ha'⟩ := ha obtain ⟨a'', ha''⟩ := (ModuleCat.epi_iff_surjective q).1 hq a' refine ⟨p a'', ?_⟩ - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [← LinearMap.comp_apply, ← ModuleCat.comp_def, H, ModuleCat.comp_def, LinearMap.comp_apply, + dsimp at ha' ⊢ + rw [← LinearMap.comp_apply, ← ModuleCat.hom_comp, H, ModuleCat.hom_comp, LinearMap.comp_apply, ha'', ha'] end Module diff --git a/Mathlib/CategoryTheory/Limits/ConcreteCategory/WithAlgebraicStructures.lean b/Mathlib/CategoryTheory/Limits/ConcreteCategory/WithAlgebraicStructures.lean index dd4c4c861c62c..1496b467c1f8a 100644 --- a/Mathlib/CategoryTheory/Limits/ConcreteCategory/WithAlgebraicStructures.lean +++ b/Mathlib/CategoryTheory/Limits/ConcreteCategory/WithAlgebraicStructures.lean @@ -4,12 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang -/ import Mathlib.CategoryTheory.Limits.ConcreteCategory.Basic -import Mathlib.Algebra.Module.LinearMap.Defs +import Mathlib.Algebra.Category.ModuleCat.Basic import Mathlib.Tactic.CategoryTheory.Elementwise /-! - -# Colimits in concrete categories with algebraic structures +# Colimits in ModuleCat Let `C` be a concrete category and `F : J ⥤ C` a filtered diagram in `C`. We discuss some results about `colimit F` when objects and morphisms in `C` have some algebraic structures. @@ -25,6 +24,11 @@ about `colimit F` when objects and morphisms in `C` have some algebraic structur and `x : Fᵢ` we have `r • x = 0` implies `x = 0`, then if `r • x = 0` for `x : colimit F`, then `x = 0`. +## Implementation details + +For now, we specialize our results to `C = ModuleCat R`, which is the only place they are used. +In the future they might be generalized by assuming a `HasForget₂ C (ModuleCat R)` instance, +plus assertions that the module structures induced by `HasForget₂` coincide. -/ universe t w v u r @@ -33,20 +37,20 @@ open CategoryTheory namespace CategoryTheory.Limits.Concrete -attribute [local instance] ConcreteCategory.instFunLike ConcreteCategory.hasCoeToSort - -variable {C : Type u} [Category.{v} C] [ConcreteCategory.{max t w} C] {J : Type w} [Category.{r} J] +variable (R : Type*) [Ring R] {J : Type w} [Category.{r} J] section zero theorem colimit_rep_eq_zero - (F : J ⥤ C) [PreservesColimit F (forget C)] [IsFiltered J] - [∀ c : C, Zero c] [∀ {c c' : C}, ZeroHomClass (c ⟶ c') c c'] [HasColimit F] - (j : J) (x : F.obj j) (hx : colimit.ι F j x = 0) : - ∃ (j' : J) (i : j ⟶ j'), F.map i x = 0 := by - rw [show 0 = colimit.ι F j 0 by simp, colimit_rep_eq_iff_exists] at hx + (F : J ⥤ ModuleCat.{max t w} R) [PreservesColimit F (forget (ModuleCat R))] [IsFiltered J] + [HasColimit F] (j : J) (x : F.obj j) (hx : colimit.ι F j x = 0) : + ∃ (j' : J) (i : j ⟶ j'), (F.map i).hom x = 0 := by + -- Break the abstraction barrier between homs and functions for `colimit_rep_eq_iff_exists`. + have : ∀ (X Y : ModuleCat R) (f : X ⟶ Y), + DFunLike.coe f.hom = DFunLike.coe (self := ConcreteCategory.instFunLike) f := fun _ _ _ => rfl + rw [show 0 = colimit.ι F j 0 by simp, this, colimit_rep_eq_iff_exists] at hx obtain ⟨j', i, y, g⟩ := hx - exact ⟨j', i, g ▸ by simp⟩ + exact ⟨j', i, g ▸ by simp [← this]⟩ end zero @@ -57,23 +61,25 @@ if `r` has no zero smul divisors for all small-enough sections, then `r` has no in the colimit. -/ lemma colimit_no_zero_smul_divisor - (F : J ⥤ C) [PreservesColimit F (forget C)] [IsFiltered J] [HasColimit F] - (R : Type*) [Semiring R] - [∀ c : C, AddCommMonoid c] [∀ c : C, Module R c] - [∀ {c c' : C}, LinearMapClass (c ⟶ c') R c c'] + (F : J ⥤ ModuleCat.{max t w} R) [PreservesColimit F (forget (ModuleCat R))] + [IsFiltered J] [HasColimit F] (r : R) (H : ∃ (j' : J), ∀ (j : J) (_ : j' ⟶ j), ∀ (c : F.obj j), r • c = 0 → c = 0) - (x : (forget C).obj (colimit F)) (hx : r • x = 0) : x = 0 := by + (x : (forget (ModuleCat R)).obj (colimit F)) (hx : r • x = 0) : x = 0 := by + + -- Break the abstraction barrier between homs and functions for `Concrete.colimit_exists_rep`. + have : ∀ (X Y : ModuleCat R) (f : X ⟶ Y), + DFunLike.coe f.hom = DFunLike.coe (self := ConcreteCategory.instFunLike) f := fun _ _ _ => rfl classical obtain ⟨j, x, rfl⟩ := Concrete.colimit_exists_rep F x - rw [← map_smul] at hx + rw [← this, ← map_smul (colimit.ι F j).hom] at hx obtain ⟨j', i, h⟩ := Concrete.colimit_rep_eq_zero (hx := hx) obtain ⟨j'', H⟩ := H - simpa [elementwise_of% (colimit.w F), map_zero] using congr(colimit.ι F _ + simpa [elementwise_of% (colimit.w F), this, map_zero] using congr(colimit.ι F _ $(H (IsFiltered.sup {j, j', j''} { ⟨j, j', by simp, by simp, i⟩ }) (IsFiltered.toSup _ _ <| by simp) (F.map (IsFiltered.toSup _ _ <| by simp) x) (by rw [← IsFiltered.toSup_commutes (f := i) (mY := by simp) (mf := by simp), F.map_comp, - comp_apply, ← map_smul, ← map_smul, h, map_zero]))) + ModuleCat.comp_apply, ← map_smul, ← map_smul, h, map_zero]))) end module diff --git a/Mathlib/CategoryTheory/Linear/Yoneda.lean b/Mathlib/CategoryTheory/Linear/Yoneda.lean index 4836d3c76a695..ae48cbcbb39b7 100644 --- a/Mathlib/CategoryTheory/Linear/Yoneda.lean +++ b/Mathlib/CategoryTheory/Linear/Yoneda.lean @@ -28,9 +28,6 @@ namespace CategoryTheory variable (R : Type w) [Ring R] {C : Type u} [Category.{v} C] [Preadditive C] [Linear R C] variable (C) --- Porting note: inserted specific `ModuleCat.ofHom` in the definition of `linearYoneda` --- and similarly in `linearCoyoneda`, otherwise many simp lemmas are not triggered automatically. --- Eventually, doing so allows more proofs to be automatic! /-- The Yoneda embedding for `R`-linear categories `C`, sending an object `X : C` to the `ModuleCat R`-valued presheaf on `C`, with value on `Y : Cᵒᵖ` given by `ModuleCat.of R (unop Y ⟶ X)`. -/ diff --git a/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean b/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean index 9341953c58518..94382ab42f817 100644 --- a/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean +++ b/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean @@ -45,19 +45,19 @@ instance Ring_of_Mon_ (A : Mon_ (ModuleCat.{u} R)) : Ring A.X := one := A.one (1 : R) mul := fun x y => A.mul (x ⊗ₜ y) one_mul := fun x => by - convert LinearMap.congr_fun A.one_mul ((1 : R) ⊗ₜ x) + convert LinearMap.congr_fun (ModuleCat.hom_ext_iff.mp A.one_mul) ((1 : R) ⊗ₜ x) rw [MonoidalCategory.leftUnitor_hom_apply, one_smul] mul_one := fun x => by - convert LinearMap.congr_fun A.mul_one (x ⊗ₜ (1 : R)) - erw [MonoidalCategory.leftUnitor_hom_apply, one_smul] + convert LinearMap.congr_fun (ModuleCat.hom_ext_iff.mp A.mul_one) (x ⊗ₜ (1 : R)) + rw [MonoidalCategory.rightUnitor_hom_apply, one_smul] mul_assoc := fun x y z => by - convert LinearMap.congr_fun A.mul_assoc (x ⊗ₜ y ⊗ₜ z) + convert LinearMap.congr_fun (ModuleCat.hom_ext_iff.mp A.mul_assoc) (x ⊗ₜ y ⊗ₜ z) left_distrib := fun x y z => by - convert A.mul.map_add (x ⊗ₜ y) (x ⊗ₜ z) + convert A.mul.hom.map_add (x ⊗ₜ y) (x ⊗ₜ z) rw [← TensorProduct.tmul_add] rfl right_distrib := fun x y z => by - convert A.mul.map_add (x ⊗ₜ z) (y ⊗ₜ z) + convert A.mul.hom.map_add (x ⊗ₜ z) (y ⊗ₜ z) rw [← TensorProduct.add_tmul] rfl zero_mul := fun x => show A.mul _ = 0 by @@ -66,18 +66,19 @@ instance Ring_of_Mon_ (A : Mon_ (ModuleCat.{u} R)) : Ring A.X := rw [TensorProduct.tmul_zero, map_zero] } instance Algebra_of_Mon_ (A : Mon_ (ModuleCat.{u} R)) : Algebra R A.X := - { A.one with - map_zero' := A.one.map_zero + { A.one.hom with + map_zero' := A.one.hom.map_zero map_one' := rfl map_mul' := fun x y => by - have h := LinearMap.congr_fun A.one_mul.symm (x ⊗ₜ A.one y) - rwa [MonoidalCategory.leftUnitor_hom_apply, ← A.one.map_smul] at h + have h := LinearMap.congr_fun (ModuleCat.hom_ext_iff.mp A.one_mul.symm) (x ⊗ₜ A.one y) + rwa [MonoidalCategory.leftUnitor_hom_apply, ← A.one.hom.map_smul] at h commutes' := fun r a => by dsimp - have h₁ := LinearMap.congr_fun A.one_mul (r ⊗ₜ a) - have h₂ := LinearMap.congr_fun A.mul_one (a ⊗ₜ r) + have h₁ := LinearMap.congr_fun (ModuleCat.hom_ext_iff.mp A.one_mul) (r ⊗ₜ a) + have h₂ := LinearMap.congr_fun (ModuleCat.hom_ext_iff.mp A.mul_one) (a ⊗ₜ r) exact h₁.trans h₂.symm - smul_def' := fun r a => (LinearMap.congr_fun A.one_mul (r ⊗ₜ a)).symm } + smul_def' := fun r a => + (LinearMap.congr_fun (ModuleCat.hom_ext_iff.mp A.one_mul) (r ⊗ₜ a)).symm } @[simp] theorem algebraMap (A : Mon_ (ModuleCat.{u} R)) (r : R) : algebraMap R A.X r = A.one r := @@ -89,33 +90,35 @@ theorem algebraMap (A : Mon_ (ModuleCat.{u} R)) (r : R) : algebraMap R A.X r = A def functor : Mon_ (ModuleCat.{u} R) ⥤ AlgebraCat R where obj A := AlgebraCat.of R A.X map {_ _} f := AlgebraCat.ofHom - { f.hom.toAddMonoidHom with + { f.hom.hom.toAddMonoidHom with toFun := f.hom - map_one' := LinearMap.congr_fun f.one_hom (1 : R) - map_mul' := fun x y => LinearMap.congr_fun f.mul_hom (x ⊗ₜ y) - commutes' := fun r => LinearMap.congr_fun f.one_hom r } + map_one' := LinearMap.congr_fun (ModuleCat.hom_ext_iff.mp f.one_hom) (1 : R) + map_mul' := fun x y => LinearMap.congr_fun (ModuleCat.hom_ext_iff.mp f.mul_hom) (x ⊗ₜ y) + commutes' := fun r => LinearMap.congr_fun (ModuleCat.hom_ext_iff.mp f.one_hom) r } /-- Converting a bundled algebra to a monoid object in `ModuleCat R`. -/ @[simps] def inverseObj (A : AlgebraCat.{u} R) : Mon_ (ModuleCat.{u} R) where X := ModuleCat.of R A - one := Algebra.linearMap R A - mul := LinearMap.mul' R A + one := ofHom <| Algebra.linearMap R A + mul := ofHom <| LinearMap.mul' R A one_mul := by + ext : 1 -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): `ext` did not pick up `TensorProduct.ext` refine TensorProduct.ext <| LinearMap.ext_ring <| LinearMap.ext fun x => ?_ - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [compr₂_apply, compr₂_apply, CategoryTheory.comp_apply] + rw [compr₂_apply, compr₂_apply, hom_comp, LinearMap.comp_apply] -- Porting note: this `dsimp` does nothing -- dsimp [AlgebraCat.id_apply, TensorProduct.mk_apply, Algebra.linearMap_apply, - -- LinearMap.compr₂_apply, Function.comp_apply, RingHom.map_one, - -- ModuleCat.MonoidalCategory.hom_apply, AlgebraCat.coe_comp, - -- ModuleCat.MonoidalCategory.leftUnitor_hom_apply] + -- LinearMap.compr₂_apply, Function.comp_apply, RingHom.map_one, + -- ModuleCat.MonoidalCategory.tensorHom_tmul, AlgebraCat.hom_comp, + -- ModuleCat.MonoidalCategory.leftUnitor_hom_apply] -- Porting note: because `dsimp` is not effective, `rw` needs to be changed to `erw` + dsimp erw [LinearMap.mul'_apply, MonoidalCategory.leftUnitor_hom_apply, ← Algebra.smul_def] - erw [id_apply] + dsimp mul_one := by + ext : 1 -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): `ext` did not pick up `TensorProduct.ext` refine TensorProduct.ext <| LinearMap.ext fun x => LinearMap.ext_ring ?_ -- Porting note: this `dsimp` does nothing @@ -124,23 +127,22 @@ def inverseObj (A : AlgebraCat.{u} R) : Mon_ (ModuleCat.{u} R) where -- AlgebraCat.coe_comp] -- Porting note: because `dsimp` is not effective, `rw` needs to be changed to `erw` erw [compr₂_apply, compr₂_apply] - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [CategoryTheory.comp_apply] + rw [ModuleCat.hom_comp, LinearMap.comp_apply] erw [LinearMap.mul'_apply, ModuleCat.MonoidalCategory.rightUnitor_hom_apply, ← Algebra.commutes, ← Algebra.smul_def] - erw [id_apply] + dsimp mul_assoc := by + ext : 1 set_option tactic.skipAssignedInstances false in -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): `ext` did not pick up `TensorProduct.ext` refine TensorProduct.ext <| TensorProduct.ext <| LinearMap.ext fun x => LinearMap.ext fun y => LinearMap.ext fun z => ?_ dsimp only [compr₂_apply, TensorProduct.mk_apply] rw [compr₂_apply, compr₂_apply] - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [CategoryTheory.comp_apply, - CategoryTheory.comp_apply, CategoryTheory.comp_apply] + rw [hom_comp, LinearMap.comp_apply, hom_comp, LinearMap.comp_apply, hom_comp, + LinearMap.comp_apply] erw [LinearMap.mul'_apply, LinearMap.mul'_apply] - erw [id_apply] + dsimp only [id_coe, id_eq] erw [TensorProduct.mk_apply, TensorProduct.mk_apply, mul'_apply, LinearMap.id_apply, mul'_apply] simp only [LinearMap.mul'_apply, mul_assoc] @@ -150,9 +152,9 @@ def inverseObj (A : AlgebraCat.{u} R) : Mon_ (ModuleCat.{u} R) where def inverse : AlgebraCat.{u} R ⥤ Mon_ (ModuleCat.{u} R) where obj := inverseObj map f := - { hom := f.hom.toLinearMap - one_hom := LinearMap.ext f.hom.commutes - mul_hom := TensorProduct.ext <| LinearMap.ext₂ <| map_mul f.hom } + { hom := ofHom <| f.hom.toLinearMap + one_hom := hom_ext <| LinearMap.ext f.hom.commutes + mul_hom := hom_ext <| TensorProduct.ext <| LinearMap.ext₂ <| map_mul f.hom } end MonModuleEquivalenceAlgebra @@ -169,21 +171,23 @@ def monModuleEquivalenceAlgebra : Mon_ (ModuleCat.{u} R) ≌ AlgebraCat R where NatIso.ofComponents (fun A => { hom := - { hom := + { hom := ofHom { toFun := _root_.id map_add' := fun _ _ => rfl map_smul' := fun _ _ => rfl } mul_hom := by + ext : 1 -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): `ext` did not pick up `TensorProduct.ext` refine TensorProduct.ext ?_ dsimp at * rfl } inv := - { hom := + { hom := ofHom { toFun := _root_.id map_add' := fun _ _ => rfl map_smul' := fun _ _ => rfl } mul_hom := by + ext : 1 -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): `ext` did not pick up `TensorProduct.ext` refine TensorProduct.ext ?_ dsimp at * @@ -214,11 +218,11 @@ def monModuleEquivalenceAlgebraForget : Mon_.forget (ModuleCat.{u} R) := NatIso.ofComponents (fun A => - { hom := + { hom := ofHom { toFun := _root_.id map_add' := fun _ _ => rfl map_smul' := fun _ _ => rfl } - inv := + inv := ofHom { toFun := _root_.id map_add' := fun _ _ => rfl map_smul' := fun _ _ => rfl } }) diff --git a/Mathlib/Condensed/Discrete/Module.lean b/Mathlib/Condensed/Discrete/Module.lean index 5657c6d51579e..75760ce80c094 100644 --- a/Mathlib/Condensed/Discrete/Module.lean +++ b/Mathlib/Condensed/Discrete/Module.lean @@ -40,8 +40,8 @@ constant maps. def functorToPresheaves : ModuleCat.{max u w} R ⥤ ((CompHausLike.{u} P)ᵒᵖ ⥤ ModuleCat R) where obj X := { obj := fun ⟨S⟩ ↦ ModuleCat.of R (LocallyConstant S X) - map := fun f ↦ comapₗ R f.unop } - map f := { app := fun S ↦ mapₗ R f } + map := fun f ↦ ModuleCat.ofHom (comapₗ R f.unop) } + map f := { app := fun S ↦ ModuleCat.ofHom (mapₗ R f.hom) } variable [HasExplicitFiniteCoproducts.{0} P] [HasExplicitPullbacks.{u} P] (hs : ∀ ⦃X Y : CompHausLike P⦄ (f : X ⟶ Y), EffectiveEpi f → Function.Surjective f) @@ -79,8 +79,8 @@ abbrev functor : ModuleCat R ⥤ CondensedMod.{u} R := /-- Auxiliary definition for `functorIsoDiscrete`. -/ noncomputable def functorIsoDiscreteAux₁ (M : ModuleCat.{u+1} R) : M ≅ (ModuleCat.of R (LocallyConstant (CompHaus.of PUnit.{u+1}) M)) where - hom := constₗ R - inv := evalₗ R PUnit.unit + hom := ModuleCat.ofHom (constₗ R) + inv := ModuleCat.ofHom (evalₗ R PUnit.unit) /-- Auxiliary definition for `functorIsoDiscrete`. -/ noncomputable def functorIsoDiscreteAux₂ (M : ModuleCat R) : @@ -186,8 +186,8 @@ abbrev functor : ModuleCat R ⥤ LightCondMod.{u} R := /-- Auxiliary definition for `functorIsoDiscrete`. -/ noncomputable def functorIsoDiscreteAux₁ (M : ModuleCat.{u} R) : M ≅ (ModuleCat.of R (LocallyConstant (LightProfinite.of PUnit.{u+1}) M)) where - hom := constₗ R - inv := evalₗ R PUnit.unit + hom := ModuleCat.ofHom (constₗ R) + inv := ModuleCat.ofHom (evalₗ R PUnit.unit) /-- Auxiliary definition for `functorIsoDiscrete`. -/ noncomputable def functorIsoDiscreteAux₂ (M : ModuleCat.{u} R) : diff --git a/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat.lean b/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat.lean index 7e6c8fa28492e..5d6cf88a2bfcf 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat.lean @@ -37,8 +37,9 @@ instance : CoeSort (QuadraticModuleCat.{v} R) (Type v) := /-- The object in the category of quadratic R-modules associated to a quadratic R-module. -/ @[simps form] def of {X : Type v} [AddCommGroup X] [Module R X] (Q : QuadraticForm R X) : - QuadraticModuleCat R where - form := Q + QuadraticModuleCat R := + { ModuleCat.of R X with + form := Q } /-- A type alias for `QuadraticForm.LinearIsometry` to avoid confusion between the categorical and algebraic spellings of composition. -/ @@ -89,7 +90,7 @@ instance concreteCategory : ConcreteCategory.{v} (QuadraticModuleCat.{v} R) wher instance hasForgetToModule : HasForget₂ (QuadraticModuleCat R) (ModuleCat R) where forget₂ := { obj := fun M => ModuleCat.of R M - map := fun f => f.toIsometry.toLinearMap } + map := fun f => ModuleCat.ofHom f.toIsometry.toLinearMap } @[simp] theorem forget₂_obj (X : QuadraticModuleCat R) : @@ -98,7 +99,8 @@ theorem forget₂_obj (X : QuadraticModuleCat R) : @[simp] theorem forget₂_map (X Y : QuadraticModuleCat R) (f : X ⟶ Y) : - (forget₂ (QuadraticModuleCat R) (ModuleCat R)).map f = f.toIsometry.toLinearMap := + (forget₂ (QuadraticModuleCat R) (ModuleCat R)).map f = + ModuleCat.ofHom f.toIsometry.toLinearMap := rfl variable {X Y Z : Type v} diff --git a/Mathlib/RepresentationTheory/Character.lean b/Mathlib/RepresentationTheory/Character.lean index 5479f9417977f..cb98ce35c607a 100644 --- a/Mathlib/RepresentationTheory/Character.lean +++ b/Mathlib/RepresentationTheory/Character.lean @@ -116,8 +116,8 @@ theorem char_orthonormal (V W : FDRep k G) [Simple V] [Simple W] : rw [char_iso (FDRep.dualTensorIsoLinHom W.ρ V)] -- The average over the group of the character of a representation equals the dimension of the -- space of invariants. - rw [average_char_eq_finrank_invariants] - rw [show (of (linHom W.ρ V.ρ)).ρ = linHom W.ρ V.ρ from FDRep.of_ρ (linHom W.ρ V.ρ)] + rw [average_char_eq_finrank_invariants, ← FDRep.endMulEquiv_comp_ρ (of _), + FDRep.of_ρ (linHom W.ρ V.ρ)] -- The space of invariants of `Hom(W, V)` is the subspace of `G`-equivariant linear maps, -- `Hom_G(W, V)`. erw [(linHom.invariantsEquivFDRepHom W V).finrank_eq] -- Porting note: Changed `rw` to `erw` diff --git a/Mathlib/RepresentationTheory/FDRep.lean b/Mathlib/RepresentationTheory/FDRep.lean index 1f3f8535f0114..c3217c02d5c69 100644 --- a/Mathlib/RepresentationTheory/FDRep.lean +++ b/Mathlib/RepresentationTheory/FDRep.lean @@ -82,7 +82,18 @@ instance (V W : FDRep k G) : FiniteDimensional k (V ⟶ W) := /-- The monoid homomorphism corresponding to the action of `G` onto `V : FDRep k G`. -/ def ρ (V : FDRep k G) : G →* V →ₗ[k] V := - Action.ρ V + (ModuleCat.endMulEquiv _).toMonoidHom.comp (Action.ρ V) + +@[simp] +lemma endMulEquiv_symm_comp_ρ (V : FDRep k G) : + (MonoidHomClass.toMonoidHom (ModuleCat.endMulEquiv V.V.obj).symm).comp (ρ V) = Action.ρ V := rfl + +@[simp] +lemma endMulEquiv_comp_ρ (V : FDRep k G) : + (MonoidHomClass.toMonoidHom (ModuleCat.endMulEquiv V.V.obj)).comp (Action.ρ V) = ρ V := rfl + +@[simp] +lemma hom_action_ρ (V : FDRep k G) (g : G) : (Action.ρ V g).hom = ρ V g := rfl /-- The underlying `LinearEquiv` of an isomorphism of representations. -/ def isoToLinearEquiv {V W : FDRep k G} (i : V ≅ W) : V ≃ₗ[k] W := @@ -91,15 +102,16 @@ def isoToLinearEquiv {V W : FDRep k G} (i : V ≅ W) : V ≃ₗ[k] W := theorem Iso.conj_ρ {V W : FDRep k G} (i : V ≅ W) (g : G) : W.ρ g = (FDRep.isoToLinearEquiv i).conj (V.ρ g) := by -- Porting note: Changed `rw` to `erw` - erw [FDRep.isoToLinearEquiv, ← FGModuleCat.Iso.conj_eq_conj, Iso.conj_apply] - rw [Iso.eq_inv_comp ((Action.forget (FGModuleCat k) (MonCat.of G)).mapIso i)] + erw [FDRep.isoToLinearEquiv, ← hom_action_ρ V, ← FGModuleCat.Iso.conj_hom_eq_conj, Iso.conj_apply] + rw [← ModuleCat.hom_ofHom (W.ρ g), ← ModuleCat.hom_ext_iff, + Iso.eq_inv_comp ((Action.forget (FGModuleCat k) (MonCat.of G)).mapIso i)] exact (i.hom.comm g).symm /-- Lift an unbundled representation to `FDRep`. -/ @[simps ρ] def of {V : Type u} [AddCommGroup V] [Module k V] [FiniteDimensional k V] (ρ : Representation k G V) : FDRep k G := - ⟨FGModuleCat.of k V, ρ⟩ + ⟨FGModuleCat.of k V, ρ ≫ MonCat.ofHom (ModuleCat.endMulEquiv _).symm.toMonoidHom⟩ instance : HasForget₂ (FDRep k G) (Rep k G) where forget₂ := (forget₂ (FGModuleCat k) (ModuleCat k)).mapAction (MonCat.of G) @@ -179,11 +191,13 @@ noncomputable def dualTensorIsoLinHomAux : /-- When `V` and `W` are finite dimensional representations of a group `G`, the isomorphism `dualTensorHomEquiv k V W` of vector spaces induces an isomorphism of representations. -/ noncomputable def dualTensorIsoLinHom : FDRep.of ρV.dual ⊗ W ≅ FDRep.of (linHom ρV W.ρ) := by - refine Action.mkIso (dualTensorIsoLinHomAux ρV W) ?_ - convert dualTensorHom_comm ρV W.ρ + refine Action.mkIso (dualTensorIsoLinHomAux ρV W) (fun g => ?_) + ext : 1 + exact dualTensorHom_comm ρV W.ρ g @[simp] -theorem dualTensorIsoLinHom_hom_hom : (dualTensorIsoLinHom ρV W).hom.hom = dualTensorHom k V W := +theorem dualTensorIsoLinHom_hom_hom : + (dualTensorIsoLinHom ρV W).hom.hom = ModuleCat.ofHom (dualTensorHom k V W) := rfl end FDRep diff --git a/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean b/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean index 2b6bd6ab54337..756771dc9da34 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean @@ -121,9 +121,9 @@ morphisms in `Rep k G`) commutes with the differentials in the complex of inhomo and the homogeneous `linearYonedaObjResolution`. -/ @[nolint checkType] theorem d_eq : d n A = - (diagonalHomEquiv n A).toModuleIso.inv ≫ + ((diagonalHomEquiv n A).toModuleIso.inv ≫ (linearYonedaObjResolution A).d n (n + 1) ≫ - (diagonalHomEquiv (n + 1) A).toModuleIso.hom := by + (diagonalHomEquiv (n + 1) A).toModuleIso.hom).hom := by ext f g /- Porting note (https://github.com/leanprover-community/mathlib4/issues/11039): broken proof was simp only [ModuleCat.coe_comp, LinearEquiv.coe_coe, Function.comp_apply, @@ -146,7 +146,7 @@ and the homogeneous `linearYonedaObjResolution`. -/ -- https://github.com/leanprover-community/mathlib4/issues/5164 change d n A f g = diagonalHomEquiv (n + 1) A ((resolution k G).d (n + 1) n ≫ (diagonalHomEquiv n A).symm f) g - rw [diagonalHomEquiv_apply, Action.comp_hom, ModuleCat.comp_def, LinearMap.comp_apply, + rw [diagonalHomEquiv_apply, Action.comp_hom, ModuleCat.hom_comp, LinearMap.comp_apply, resolution.d_eq] erw [resolution.d_of (Fin.partialProd g)] simp only [map_sum, ← Finsupp.smul_single_one _ ((-1 : k) ^ _)] @@ -175,7 +175,7 @@ $$0 \to \mathrm{Fun}(G^0, A) \to \mathrm{Fun}(G^1, A) \to \mathrm{Fun}(G^2, A) \ which calculates the group cohomology of `A`. -/ noncomputable abbrev inhomogeneousCochains : CochainComplex (ModuleCat k) ℕ := CochainComplex.of (fun n => ModuleCat.of k ((Fin n → G) → A)) - (fun n => inhomogeneousCochains.d n A) fun n => by + (fun n => ModuleCat.ofHom (inhomogeneousCochains.d n A)) fun n => by /- Porting note (https://github.com/leanprover-community/mathlib4/issues/11039): broken proof was ext x y have := LinearMap.ext_iff.1 ((linearYonedaObjResolution A).d_comp_d n (n + 1) (n + 2)) @@ -184,20 +184,23 @@ noncomputable abbrev inhomogeneousCochains : CochainComplex (ModuleCat k) ℕ := LinearEquiv.toModuleIso_inv, LinearEquiv.coe_coe, LinearEquiv.symm_apply_apply, this, LinearMap.zero_apply, map_zero, Pi.zero_apply] -/ ext x - have := LinearMap.ext_iff.1 ((linearYonedaObjResolution A).d_comp_d n (n + 1) (n + 2)) - simp only [ModuleCat.comp_def, LinearMap.comp_apply] at this + have : ∀ x, _ = (0 : _ →ₗ[_] _) x := LinearMap.ext_iff.1 (ModuleCat.hom_ext_iff.mp + ((linearYonedaObjResolution A).d_comp_d n (n + 1) (n + 2))) + simp only [ModuleCat.hom_comp, LinearMap.comp_apply] at this dsimp only - simp only [d_eq, LinearEquiv.toModuleIso_inv, LinearEquiv.toModuleIso_hom, ModuleCat.coe_comp, - Function.comp_apply] + simp only [d_eq, LinearEquiv.toModuleIso_inv_hom, LinearEquiv.toModuleIso_hom_hom, + ModuleCat.hom_comp, LinearMap.comp_apply, LinearEquiv.coe_coe, ModuleCat.hom_zero] /- Porting note: I can see I need to rewrite `LinearEquiv.coe_coe` twice to at least reduce the need for `symm_apply_apply` to be an `erw`. However, even `erw` refuses to rewrite the second `coe_coe`... -/ erw [LinearEquiv.symm_apply_apply, this] - exact map_zero _ + simp only [LinearMap.zero_apply, ChainComplex.linearYonedaObj_X, linearYoneda_obj_obj_carrier, + map_zero, Pi.zero_apply, LinearMap.zero_apply] + rfl @[simp] theorem inhomogeneousCochains.d_def (n : ℕ) : - (inhomogeneousCochains A).d n (n + 1) = inhomogeneousCochains.d n A := + (inhomogeneousCochains A).d n (n + 1) = ModuleCat.ofHom (inhomogeneousCochains.d n A) := CochainComplex.of_d _ _ _ _ /-- Given a `k`-linear `G`-representation `A`, the complex of inhomogeneous cochains is isomorphic @@ -207,8 +210,15 @@ def inhomogeneousCochainsIso : inhomogeneousCochains A ≅ linearYonedaObjResolu (Rep.diagonalHomEquiv i A).toModuleIso.symm) ?_ rintro i j (h : i + 1 = j) subst h - simp only [CochainComplex.of_d, d_eq, Category.assoc, Iso.symm_hom, Iso.hom_inv_id, - Category.comp_id] + ext + simp only [ChainComplex.linearYonedaObj_X, linearYoneda_obj_obj_carrier, CochainComplex.of_x, + linearYoneda_obj_obj_isAddCommGroup, linearYoneda_obj_obj_isModule, Iso.symm_hom, + ChainComplex.linearYonedaObj_d, ModuleCat.hom_comp, linearYoneda_obj_map_hom, + Quiver.Hom.unop_op, LinearEquiv.toModuleIso_inv_hom, LinearMap.coe_comp, Function.comp_apply, + Linear.leftComp_apply, inhomogeneousCochains.d_def, d_eq, LinearEquiv.toModuleIso_hom_hom, + ModuleCat.ofHom_comp, Category.assoc, LinearEquiv.comp_coe, LinearEquiv.self_trans_symm, + LinearEquiv.refl_toLinearMap, LinearMap.id_comp, LinearEquiv.coe_coe] + rfl /-- The `n`-cocycles `Zⁿ(G, A)` of a `k`-linear `G`-representation `A`, i.e. the kernel of the `n`th differential in the complex of inhomogeneous cochains. -/ diff --git a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean index 6f7158e455f66..ff50887857737 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean @@ -142,7 +142,7 @@ square commutes: where the vertical arrows are `zeroCochainsLequiv` and `oneCochainsLequiv` respectively. -/ theorem dZero_comp_eq : dZero A ∘ₗ (zeroCochainsLequiv A) = - oneCochainsLequiv A ∘ₗ (inhomogeneousCochains A).d 0 1 := by + oneCochainsLequiv A ∘ₗ ((inhomogeneousCochains A).d 0 1).hom := by ext x y show A.ρ y (x default) - x default = _ + ({0} : Finset _).sum _ simp_rw [Fin.val_eq_zero, zero_add, pow_one, neg_smul, one_smul, @@ -163,7 +163,7 @@ square commutes: where the vertical arrows are `oneCochainsLequiv` and `twoCochainsLequiv` respectively. -/ theorem dOne_comp_eq : dOne A ∘ₗ oneCochainsLequiv A = - twoCochainsLequiv A ∘ₗ (inhomogeneousCochains A).d 1 2 := by + twoCochainsLequiv A ∘ₗ ((inhomogeneousCochains A).d 1 2).hom := by ext x y show A.ρ y.1 (x _) - x _ + x _ = _ + _ rw [Fin.sum_univ_two] @@ -185,7 +185,8 @@ square commutes: where the vertical arrows are `twoCochainsLequiv` and `threeCochainsLequiv` respectively. -/ theorem dTwo_comp_eq : - dTwo A ∘ₗ twoCochainsLequiv A = threeCochainsLequiv A ∘ₗ (inhomogeneousCochains A).d 2 3 := by + dTwo A ∘ₗ twoCochainsLequiv A = + threeCochainsLequiv A ∘ₗ ((inhomogeneousCochains A).d 2 3).hom := by ext x y show A.ρ y.1 (x _) - x _ + x _ - x _ = _ + _ dsimp @@ -201,12 +202,13 @@ theorem dOne_comp_dZero : dOne A ∘ₗ dZero A = 0 := by rfl theorem dTwo_comp_dOne : dTwo A ∘ₗ dOne A = 0 := by - show ModuleCat.ofHom (dOne A) ≫ ModuleCat.ofHom (dTwo A) = _ - have h1 : _ ≫ ModuleCat.ofHom (dOne A) = _ ≫ _ := congr_arg ModuleCat.ofHom (dOne_comp_eq A) - have h2 : _ ≫ ModuleCat.ofHom (dTwo A) = _ ≫ _ := congr_arg ModuleCat.ofHom (dTwo_comp_eq A) - simp only [← LinearEquiv.toModuleIso_hom] at h1 h2 - simp only [(Iso.eq_inv_comp _).2 h2, (Iso.eq_inv_comp _).2 h1, - Category.assoc, Iso.hom_inv_id_assoc, HomologicalComplex.d_comp_d_assoc, zero_comp, comp_zero] + show (ModuleCat.ofHom (dOne A) ≫ ModuleCat.ofHom (dTwo A)).hom = _ + have h1 := congr_arg ModuleCat.ofHom (dOne_comp_eq A) + have h2 := congr_arg ModuleCat.ofHom (dTwo_comp_eq A) + simp only [ModuleCat.ofHom_comp, ModuleCat.ofHom_comp, ← LinearEquiv.toModuleIso_hom_hom] at h1 h2 + simp only [(Iso.eq_inv_comp _).2 h2, (Iso.eq_inv_comp _).2 h1, ModuleCat.ofHom_hom, + ModuleCat.hom_ofHom, Category.assoc, Iso.hom_inv_id_assoc, HomologicalComplex.d_comp_d_assoc, + zero_comp, comp_zero, ModuleCat.hom_zero] end Differentials @@ -751,7 +753,7 @@ lemma shortComplexH0_exact : (shortComplexH0 A).Exact := by def dZeroArrowIso : Arrow.mk ((inhomogeneousCochains A).d 0 1) ≅ Arrow.mk (ModuleCat.ofHom (dZero A)) := Arrow.isoMk (zeroCochainsLequiv A).toModuleIso - (oneCochainsLequiv A).toModuleIso (dZero_comp_eq A) + (oneCochainsLequiv A).toModuleIso (ModuleCat.hom_ext (dZero_comp_eq A)) /-- The 0-cocycles of the complex of inhomogeneous cochains of `A` are isomorphic to `A.ρ.invariants`, which is a simpler type. -/ @@ -761,7 +763,7 @@ def isoZeroCocycles : cocycles A 0 ≅ ModuleCat.of k A.ρ.invariants := (dZeroArrowIso A) lemma isoZeroCocycles_hom_comp_subtype : - (isoZeroCocycles A).hom ≫ A.ρ.invariants.subtype = + (isoZeroCocycles A).hom ≫ ModuleCat.ofHom A.ρ.invariants.subtype = iCocycles A 0 ≫ (zeroCochainsLequiv A).toModuleIso.hom := by dsimp [isoZeroCocycles] apply KernelFork.mapOfIsLimit_ι @@ -788,7 +790,9 @@ short complex associated to the complex of inhomogeneous cochains of `A`. -/ @[simps! hom inv] def shortComplexH1Iso : (inhomogeneousCochains A).sc' 0 1 2 ≅ shortComplexH1 A := isoMk (zeroCochainsLequiv A).toModuleIso (oneCochainsLequiv A).toModuleIso - (twoCochainsLequiv A).toModuleIso (dZero_comp_eq A) (dOne_comp_eq A) + (twoCochainsLequiv A).toModuleIso + (ModuleCat.hom_ext (dZero_comp_eq A)) + (ModuleCat.hom_ext (dOne_comp_eq A)) /-- The 1-cocycles of the complex of inhomogeneous cochains of `A` are isomorphic to `oneCocycles A`, which is a simpler type. -/ @@ -809,7 +813,6 @@ lemma toCocycles_comp_isoOneCocycles_hom : (zeroCochainsLequiv A).toModuleIso.hom ≫ ModuleCat.ofHom (shortComplexH1 A).moduleCatToCycles := by simp [isoOneCocycles] - rfl /-- The 1st group cohomology of `A`, defined as the 1st cohomology of the complex of inhomogeneous cochains, is isomorphic to `oneCocycles A ⧸ oneCoboundaries A`, which is a simpler type. -/ @@ -836,7 +839,9 @@ isomorphic to the 2nd short complex associated to the complex of inhomogeneous c def shortComplexH2Iso : (inhomogeneousCochains A).sc' 1 2 3 ≅ shortComplexH2 A := isoMk (oneCochainsLequiv A).toModuleIso (twoCochainsLequiv A).toModuleIso - (threeCochainsLequiv A).toModuleIso (dOne_comp_eq A) (dTwo_comp_eq A) + (threeCochainsLequiv A).toModuleIso + (ModuleCat.hom_ext (dOne_comp_eq A)) + (ModuleCat.hom_ext (dTwo_comp_eq A)) /-- The 2-cocycles of the complex of inhomogeneous cochains of `A` are isomorphic to `twoCocycles A`, which is a simpler type. -/ @@ -857,7 +862,6 @@ lemma toCocycles_comp_isoTwoCocycles_hom : (oneCochainsLequiv A).toModuleIso.hom ≫ ModuleCat.ofHom (shortComplexH2 A).moduleCatToCycles := by simp [isoTwoCocycles] - rfl /-- The 2nd group cohomology of `A`, defined as the 2nd cohomology of the complex of inhomogeneous cochains, is isomorphic to `twoCocycles A ⧸ twoCoboundaries A`, which is a simpler type. -/ diff --git a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean index 564819a98563f..fedc31c3c9a40 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean @@ -320,11 +320,8 @@ theorem diagonalHomEquiv_symm_apply (f : (Fin n → G) → A) (x : Fin (n + 1) one_smul, Rep.of_ρ, Rep.Action_ρ_eq_ρ, Rep.trivial_def (x 0)⁻¹, Finsupp.llift_apply A k k] -/ simp only [LinearEquiv.trans_symm, LinearEquiv.symm_symm, LinearEquiv.trans_apply, leftRegularHomEquiv_symm_apply, Linear.homCongr_symm_apply, Iso.trans_hom, Iso.refl_inv, - Category.comp_id, Action.comp_hom, MonoidalClosed.linearHomEquivComm_symm_hom] - -- Porting note: This is a sure sign that coercions for morphisms in `ModuleCat` - -- are still not set up properly. - rw [ModuleCat.coe_comp] - simp only [ModuleCat.coe_comp, Function.comp_apply] + Category.comp_id, Action.comp_hom, MonoidalClosed.linearHomEquivComm_symm_hom, + ModuleCat.hom_comp, LinearMap.comp_apply] -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 erw [diagonalSucc_hom_single] -- The prototype linter that checks if `erw` could be replaced with `rw` would time out @@ -482,7 +479,9 @@ instance x_projective (G : Type u) [Group G] (n : ℕ) : /-- Simpler expression for the differential in the standard resolution of `k` as a `G`-representation. It sends `(g₀, ..., gₙ₊₁) ↦ ∑ (-1)ⁱ • (g₀, ..., ĝᵢ, ..., gₙ₊₁)`. -/ -theorem d_eq (n : ℕ) : ((groupCohomology.resolution k G).d (n + 1) n).hom = d k G (n + 1) := by +theorem d_eq (n : ℕ) : ((groupCohomology.resolution k G).d (n + 1) n).hom = + ModuleCat.ofHom (d k G (n + 1)) := by + ext : 1 refine Finsupp.lhom_ext' fun x => LinearMap.ext_ring ?_ dsimp [groupCohomology.resolution] /- Porting note (https://github.com/leanprover-community/mathlib4/issues/11039): broken proof was @@ -490,11 +489,12 @@ theorem d_eq (n : ℕ) : ((groupCohomology.resolution k G).d (n + 1) n).hom = d simp_rw [alternatingFaceMapComplex_obj_d, AlternatingFaceMapComplex.objD, SimplicialObject.δ, Functor.comp_map, ← Int.cast_smul_eq_zsmul k ((-1) ^ _ : ℤ), Int.cast_pow, Int.cast_neg, Int.cast_one, Action.sum_hom, Action.smul_hom, Rep.linearization_map_hom] - rw [LinearMap.coeFn_sum, Fintype.sum_apply] + rw [ModuleCat.hom_sum, LinearMap.coeFn_sum, Fintype.sum_apply] erw [d_of (k := k) x] /- Porting note: want to rewrite `LinearMap.smul_apply` but simp/simp_rw won't do it; I need erw, so using Finset.sum_congr to get rid of the binder -/ refine Finset.sum_congr rfl fun _ _ => ?_ + simp only [ModuleCat.hom_smul, SimplexCategory.len_mk] erw [LinearMap.smul_apply] rw [Finsupp.lmapDomain_apply, Finsupp.mapDomain_single, Finsupp.smul_single', mul_one] rfl @@ -533,8 +533,8 @@ def forget₂ToModuleCatHomotopyEquiv : /-- The hom of `k`-linear `G`-representations `k[G¹] → k` sending `∑ nᵢgᵢ ↦ ∑ nᵢ`. -/ def ε : Rep.ofMulAction k G (Fin 1 → G) ⟶ Rep.trivial k G k where - hom := Finsupp.linearCombination _ fun _ => (1 : k) - comm g := Finsupp.lhom_ext' fun _ => LinearMap.ext_ring (by + hom := ModuleCat.ofHom <| Finsupp.linearCombination _ fun _ => (1 : k) + comm g := ModuleCat.hom_ext <| Finsupp.lhom_ext' fun _ => LinearMap.ext_ring (by show Finsupp.linearCombination k (fun _ => (1 : k)) (Finsupp.mapDomain _ (Finsupp.single _ _)) = Finsupp.linearCombination k (fun _ => (1 : k)) (Finsupp.single _ _) @@ -551,7 +551,8 @@ theorem forget₂ToModuleCatHomotopyEquiv_f_0_eq : convert Category.id_comp (X := (forget₂ToModuleCat k G).X 0) _ · dsimp only [HomotopyEquiv.ofIso, compForgetAugmentedIso] simp only [Iso.symm_hom, eqToIso.inv, HomologicalComplex.eqToHom_f, eqToHom_refl] - trans (linearCombination _ fun _ => (1 : k)).comp ((ModuleCat.free k).map (terminal.from _)) + ext : 1 + trans (linearCombination _ fun _ => (1 : k)).comp ((ModuleCat.free k).map (terminal.from _)).hom · erw [Finsupp.lmapDomain_linearCombination (α := Fin 1 → G) (R := k) (α' := ⊤_ Type u) (v := fun _ => (1 : k)) (v' := fun _ => (1 : k)) (terminal.from @@ -563,21 +564,19 @@ theorem forget₂ToModuleCatHomotopyEquiv_f_0_eq : · ext x dsimp (config := { unfoldPartialApp := true }) [HomotopyEquiv.ofIso, Finsupp.LinearEquiv.finsuppUnique] - rw [linearCombination_single, one_smul, - @Unique.eq_default _ Types.terminalIso.toEquiv.unique x, - ChainComplex.single₀_map_f_zero, LinearMap.coe_mk, AddHom.coe_mk, Function.comp_apply, - Finsupp.equivFunOnFinite_apply, Finsupp.single_eq_same] + rw [@Unique.eq_default _ Types.terminalIso.toEquiv.unique x] + simp · exact @Subsingleton.elim _ (@Unique.instSubsingleton _ (Limits.uniqueToTerminal _)) _ _ theorem d_comp_ε : (groupCohomology.resolution k G).d 1 0 ≫ ε k G = 0 := by ext : 1 - refine LinearMap.ext fun x => ?_ + refine ModuleCat.hom_ext <| LinearMap.ext fun x => ?_ have : (forget₂ToModuleCat k G).d 1 0 ≫ (forget₂ (Rep k G) (ModuleCat.{u} k)).map (ε k G) = 0 := by rw [← forget₂ToModuleCatHomotopyEquiv_f_0_eq, ← (forget₂ToModuleCatHomotopyEquiv k G).1.2 1 0 rfl] exact comp_zero - exact LinearMap.ext_iff.1 this _ + exact LinearMap.ext_iff.1 (ModuleCat.hom_ext_iff.mp this) _ /-- The chain map from the standard resolution of `k` to `k[0]` given by `∑ nᵢgᵢ ↦ ∑ nᵢ` in degree zero. -/ diff --git a/Mathlib/RepresentationTheory/Invariants.lean b/Mathlib/RepresentationTheory/Invariants.lean index 898d8d3427513..acf708f7c06b9 100644 --- a/Mathlib/RepresentationTheory/Invariants.lean +++ b/Mathlib/RepresentationTheory/Invariants.lean @@ -121,19 +121,23 @@ variable {k : Type u} [CommRing k] {G : Grp.{u}} theorem mem_invariants_iff_comm {X Y : Rep k G} (f : X.V →ₗ[k] Y.V) (g : G) : (linHom X.ρ Y.ρ) g f = f ↔ f.comp (X.ρ g) = (Y.ρ g).comp f := by dsimp - erw [← ρAut_apply_inv] - rw [← LinearMap.comp_assoc, ← ModuleCat.comp_def, ← ModuleCat.comp_def, Iso.inv_comp_eq, - ρAut_apply_hom] + rw [← LinearMap.comp_assoc, ← ModuleCat.hom_ofHom (Y.ρ g), ← ModuleCat.hom_ofHom f, + ← ModuleCat.hom_comp, ← ModuleCat.hom_ofHom (X.ρ g⁻¹), ← ModuleCat.hom_comp, + Rep.ofHom_ρ, ← ρAut_apply_inv X g, Rep.ofHom_ρ, ← ρAut_apply_hom Y g, ← ModuleCat.hom_ext_iff, + Iso.inv_comp_eq, ρAut_apply_hom, ← ModuleCat.hom_ofHom (X.ρ g), + ← ModuleCat.hom_comp, ← ModuleCat.hom_ext_iff] exact comm /-- The invariants of the representation `linHom X.ρ Y.ρ` correspond to the representation homomorphisms from `X` to `Y`. -/ @[simps] def invariantsEquivRepHom (X Y : Rep k G) : (linHom X.ρ Y.ρ).invariants ≃ₗ[k] X ⟶ Y where - toFun f := ⟨f.val, fun g => (mem_invariants_iff_comm _ g).1 (f.property g)⟩ + toFun f := ⟨ModuleCat.ofHom f.val, fun g => + ModuleCat.hom_ext ((mem_invariants_iff_comm _ g).1 (f.property g))⟩ map_add' _ _ := rfl map_smul' _ _ := rfl - invFun f := ⟨f.hom, fun g => (mem_invariants_iff_comm _ g).2 (f.comm g)⟩ + invFun f := ⟨f.hom.hom, fun g => + (mem_invariants_iff_comm _ g).2 (ModuleCat.hom_ext_iff.mp (f.comm g))⟩ left_inv _ := by ext; rfl right_inv _ := by ext; rfl diff --git a/Mathlib/RepresentationTheory/Rep.lean b/Mathlib/RepresentationTheory/Rep.lean index e445fe7584092..342bac0e56cf4 100644 --- a/Mathlib/RepresentationTheory/Rep.lean +++ b/Mathlib/RepresentationTheory/Rep.lean @@ -61,11 +61,11 @@ instance (V : Rep k G) : Module k V := by -/ def ρ (V : Rep k G) : Representation k G V := -- Porting note: was `V.ρ` - Action.ρ V + (ModuleCat.endMulEquiv V.V).toMonoidHom.comp (Action.ρ V) /-- Lift an unbundled representation to `Rep`. -/ def of {V : Type u} [AddCommGroup V] [Module k V] (ρ : G →* V →ₗ[k] V) : Rep k G := - ⟨ModuleCat.of k V, ρ⟩ + ⟨ModuleCat.of k V, MonCat.ofHom ((ModuleCat.endMulEquiv _).symm.toMonoidHom.comp ρ) ⟩ @[simp] theorem coe_of {V : Type u} [AddCommGroup V] [Module k V] (ρ : G →* V →ₗ[k] V) : @@ -76,9 +76,16 @@ theorem coe_of {V : Type u} [AddCommGroup V] [Module k V] (ρ : G →* V →ₗ[ theorem of_ρ {V : Type u} [AddCommGroup V] [Module k V] (ρ : G →* V →ₗ[k] V) : (of ρ).ρ = ρ := rfl -theorem Action_ρ_eq_ρ {A : Rep k G} : Action.ρ A = A.ρ := +theorem Action_ρ_eq_ρ {A : Rep k G} : + Action.ρ A = (ModuleCat.endMulEquiv _).symm.toMonoidHom.comp A.ρ := rfl +@[simp] +lemma ρ_hom {X : Rep k G} (g : G) : (Action.ρ X g).hom = X.ρ g := rfl + +@[simp] +lemma ofHom_ρ {X : Rep k G} (g : G) : ModuleCat.ofHom (X.ρ g) = Action.ρ X g := rfl + /-- Allows us to apply lemmas about the underlying `ρ`, which would take an element `g : G` rather than `g : MonCat.of G` as an argument. -/ theorem of_ρ_apply {V : Type u} [AddCommGroup V] [Module k V] (ρ : Representation k G V) @@ -97,7 +104,7 @@ theorem ρ_self_inv_apply {G : Type u} [Group G] {A : Rep k G} (g : G) (x : A) : theorem hom_comm_apply {A B : Rep k G} (f : A ⟶ B) (g : G) (x : A) : f.hom (A.ρ g x) = B.ρ g (f.hom x) := - LinearMap.ext_iff.1 (f.comm g) x + LinearMap.ext_iff.1 (ModuleCat.hom_ext_iff.mp (f.comm g)) x variable (k G) @@ -172,7 +179,8 @@ theorem linearization_single (X : Action (Type u) (MonCat.of G)) (g : G) (x : X. variable {X Y : Action (Type u) (MonCat.of G)} (f : X ⟶ Y) @[simp] -theorem linearization_map_hom : ((linearization k G).map f).hom = Finsupp.lmapDomain k k f.hom := +theorem linearization_map_hom : ((linearization k G).map f).hom = + ModuleCat.ofHom (Finsupp.lmapDomain k k f.hom) := rfl theorem linearization_map_hom_single (x : X.V) (r : k) : @@ -183,16 +191,19 @@ open Functor.LaxMonoidal Functor.OplaxMonoidal Functor.Monoidal @[simp] theorem linearization_μ_hom (X Y : Action (Type u) (MonCat.of G)) : - (μ (linearization k G) X Y).hom = (finsuppTensorFinsupp' k X.V Y.V).toLinearMap := + (μ (linearization k G) X Y).hom = + ModuleCat.ofHom (finsuppTensorFinsupp' k X.V Y.V).toLinearMap := rfl @[simp] theorem linearization_δ_hom (X Y : Action (Type u) (MonCat.of G)) : - (δ (linearization k G) X Y).hom = (finsuppTensorFinsupp' k X.V Y.V).symm.toLinearMap := + (δ (linearization k G) X Y).hom = + ModuleCat.ofHom (finsuppTensorFinsupp' k X.V Y.V).symm.toLinearMap := rfl @[simp] -theorem linearization_ε_hom : (ε (linearization k G)).hom = Finsupp.lsingle PUnit.unit := +theorem linearization_ε_hom : (ε (linearization k G)).hom = + ModuleCat.ofHom (Finsupp.lsingle PUnit.unit) := rfl theorem linearization_η_hom_apply (r : k) : @@ -206,7 +217,7 @@ on `k[X]`. -/ @[simps!] noncomputable def linearizationTrivialIso (X : Type u) : (linearization k G).obj (Action.mk X 1) ≅ trivial k G (X →₀ k) := - Action.mkIso (Iso.refl _) fun _ => Finsupp.lhom_ext' fun _ => LinearMap.ext + Action.mkIso (Iso.refl _) fun _ => ModuleCat.hom_ext <| Finsupp.lhom_ext' fun _ => LinearMap.ext fun _ => linearization_single .. /-- Given a `G`-action on `H`, this is `k[H]` bundled with the natural representation @@ -269,14 +280,15 @@ variable {k G} `g ↦ A.ρ(g)(x).` -/ @[simps] noncomputable def leftRegularHom (A : Rep k G) (x : A) : Rep.ofMulAction k G G ⟶ A where - hom := Finsupp.lift _ _ _ fun g => A.ρ g x + hom := ModuleCat.ofHom <| Finsupp.lift _ _ _ fun g => A.ρ g x comm g := by + ext : 1 refine Finsupp.lhom_ext' fun y => LinearMap.ext_ring ?_ /- Porting note: rest of broken proof was simpa only [LinearMap.comp_apply, ModuleCat.comp_def, Finsupp.lsingle_apply, Finsupp.lift_apply, Action_ρ_eq_ρ, of_ρ_apply, Representation.ofMulAction_single, Finsupp.sum_single_index, zero_smul, one_smul, smul_eq_mul, A.ρ.map_mul] -/ - simp only [LinearMap.comp_apply, ModuleCat.comp_def, Finsupp.lsingle_apply] + simp only [LinearMap.comp_apply, ModuleCat.hom_comp, Finsupp.lsingle_apply] erw [Finsupp.lift_apply, Finsupp.lift_apply, Representation.ofMulAction_single (G := G)] simp only [Finsupp.sum_single_index, zero_smul, one_smul, smul_eq_mul, A.ρ.map_mul, of_ρ] rfl @@ -284,7 +296,7 @@ noncomputable def leftRegularHom (A : Rep k G) (x : A) : Rep.ofMulAction k G G theorem leftRegularHom_apply {A : Rep k G} (x : A) : (leftRegularHom A x).hom (Finsupp.single 1 1) = x := by -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [leftRegularHom_hom, Finsupp.lift_apply, Finsupp.sum_single_index, one_smul, + erw [leftRegularHom_hom_hom, Finsupp.lift_apply, Finsupp.sum_single_index, one_smul, A.ρ.map_one, LinearMap.one_apply] rw [zero_smul] @@ -297,14 +309,14 @@ noncomputable def leftRegularHomEquiv (A : Rep k G) : (Rep.ofMulAction k G G ⟶ map_smul' _ _ := rfl invFun x := leftRegularHom A x left_inv f := by - refine Action.Hom.ext (Finsupp.lhom_ext' fun x : G => LinearMap.ext_ring ?_) + refine Action.Hom.ext (ModuleCat.hom_ext (Finsupp.lhom_ext' fun x : G => LinearMap.ext_ring ?_)) have : f.hom ((ofMulAction k G G).ρ x (Finsupp.single (1 : G) (1 : k))) = A.ρ x (f.hom (Finsupp.single (1 : G) (1 : k))) := - LinearMap.ext_iff.1 (f.comm x) (Finsupp.single 1 1) - simp only [leftRegularHom_hom, LinearMap.comp_apply, Finsupp.lsingle_apply, Finsupp.lift_apply, - ← this, coe_of, of_ρ, Representation.ofMulAction_single x (1 : G) (1 : k), smul_eq_mul, - mul_one, zero_smul, Finsupp.sum_single_index, one_smul] + LinearMap.ext_iff.1 (ModuleCat.hom_ext_iff.mp (f.comm x)) (Finsupp.single 1 1) + simp only [leftRegularHom_hom_hom, LinearMap.comp_apply, Finsupp.lsingle_apply, + Finsupp.lift_apply, ← this, coe_of, of_ρ, Representation.ofMulAction_single x (1 : G) (1 : k), + smul_eq_mul, mul_one, zero_smul, Finsupp.sum_single_index, one_smul] -- Mismatched `Zero k` instances rfl right_inv x := leftRegularHom_apply x @@ -312,7 +324,7 @@ noncomputable def leftRegularHomEquiv (A : Rep k G) : (Rep.ofMulAction k G G ⟶ theorem leftRegularHomEquiv_symm_single {A : Rep k G} (x : A) (g : G) : ((leftRegularHomEquiv A).symm x).hom (Finsupp.single g 1) = A.ρ g x := by -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [leftRegularHomEquiv_symm_apply, leftRegularHom_hom, Finsupp.lift_apply, + erw [leftRegularHomEquiv_symm_apply, leftRegularHom_hom_hom, Finsupp.lift_apply, Finsupp.sum_single_index, one_smul] rw [zero_smul] @@ -332,8 +344,8 @@ variable [Group G] (A B C : Rep k G) protected def ihom (A : Rep k G) : Rep k G ⥤ Rep k G where obj B := Rep.of (Representation.linHom A.ρ B.ρ) map := fun {X} {Y} f => - { hom := ModuleCat.ofHom (LinearMap.llcomp k _ _ _ f.hom) - comm := fun g => LinearMap.ext fun x => LinearMap.ext fun y => by + { hom := ModuleCat.ofHom (LinearMap.llcomp k _ _ _ f.hom.hom) + comm := fun g => ModuleCat.hom_ext <| LinearMap.ext fun x => LinearMap.ext fun y => by show f.hom (X.ρ g _) = _ simp only [hom_comm_apply]; rfl } map_id := fun _ => by ext; rfl @@ -348,32 +360,33 @@ protected def ihom (A : Rep k G) : Rep k G ⥤ Rep k G where `k`-linear map underlying `f`, giving a map `A →ₗ[k] B →ₗ[k] C`, then flipping the arguments. -/ def homEquiv (A B C : Rep k G) : (A ⊗ B ⟶ C) ≃ (B ⟶ (Rep.ihom A).obj C) where toFun f := - { hom := (TensorProduct.curry f.hom).flip + { hom := ModuleCat.ofHom <| (TensorProduct.curry f.hom.hom).flip comm := fun g => by - refine LinearMap.ext fun x => LinearMap.ext fun y => ?_ + ext x : 2 + refine LinearMap.ext fun y => ?_ change f.hom (_ ⊗ₜ[k] _) = C.ρ g (f.hom (_ ⊗ₜ[k] _)) rw [← hom_comm_apply] change _ = f.hom ((A.ρ g * A.ρ g⁻¹) y ⊗ₜ[k] _) simp only [← map_mul, mul_inv_cancel, map_one] rfl } invFun f := - { hom := TensorProduct.uncurry k _ _ _ f.hom.flip - comm := fun g => TensorProduct.ext' fun x y => by -/- Porting note: rest of broken proof was + { hom := ModuleCat.ofHom <| TensorProduct.uncurry k _ _ _ f.hom.hom.flip + comm := fun g => ModuleCat.hom_ext <| TensorProduct.ext' fun x y => by + /- Porting note: rest of broken proof was dsimp only [MonoidalCategory.tensorLeft_obj, ModuleCat.comp_def, LinearMap.comp_apply, tensor_ρ, ModuleCat.MonoidalCategory.hom_apply, TensorProduct.map_tmul] simp only [TensorProduct.uncurry_apply f.hom.flip, LinearMap.flip_apply, Action_ρ_eq_ρ, hom_comm_apply f g y, Rep.ihom_obj_ρ_apply, LinearMap.comp_apply, ρ_inv_self_apply] -/ - change TensorProduct.uncurry k _ _ _ f.hom.flip (A.ρ g x ⊗ₜ[k] B.ρ g y) = - C.ρ g (TensorProduct.uncurry k _ _ _ f.hom.flip (x ⊗ₜ[k] y)) + change TensorProduct.uncurry k _ _ _ f.hom.hom.flip (A.ρ g x ⊗ₜ[k] B.ρ g y) = + C.ρ g (TensorProduct.uncurry k _ _ _ f.hom.hom.flip (x ⊗ₜ[k] y)) -- The next 3 tactics used to be `rw` before https://github.com/leanprover/lean4/pull/2644 erw [TensorProduct.uncurry_apply, LinearMap.flip_apply, hom_comm_apply, Rep.ihom_obj_ρ_apply, LinearMap.comp_apply, LinearMap.comp_apply] --, ρ_inv_self_apply (A := C)] dsimp rw [ρ_inv_self_apply] - rfl} - left_inv _ := Action.Hom.ext (TensorProduct.ext' fun _ _ => rfl) + rfl } + left_inv _ := Action.Hom.ext (ModuleCat.hom_ext (TensorProduct.ext' fun _ _ => rfl)) right_inv f := by ext; rfl variable {A B C} @@ -381,12 +394,13 @@ variable {A B C} /-- Porting note: if we generate this with `@[simps]` the linter complains some types in the LHS simplify. -/ theorem homEquiv_apply_hom (f : A ⊗ B ⟶ C) : - (homEquiv A B C f).hom = (TensorProduct.curry f.hom).flip := rfl + (homEquiv A B C f).hom = ModuleCat.ofHom (TensorProduct.curry f.hom.hom).flip := rfl /-- Porting note: if we generate this with `@[simps]` the linter complains some types in the LHS simplify. -/ theorem homEquiv_symm_apply_hom (f : B ⟶ (Rep.ihom A).obj C) : - ((homEquiv A B C).symm f).hom = TensorProduct.uncurry k A B C f.hom.flip := rfl + ((homEquiv A B C).symm f).hom = + ModuleCat.ofHom (TensorProduct.uncurry k A B C f.hom.hom.flip) := rfl instance : MonoidalClosed (Rep k G) where closed A := @@ -394,9 +408,9 @@ instance : MonoidalClosed (Rep k G) where adj := Adjunction.mkOfHomEquiv ( { homEquiv := Rep.homEquiv A homEquiv_naturality_left_symm := fun _ _ => Action.Hom.ext - (TensorProduct.ext' fun _ _ => rfl) - homEquiv_naturality_right := fun _ _ => Action.Hom.ext (LinearMap.ext - fun _ => LinearMap.ext fun _ => rfl) })} + (ModuleCat.hom_ext (TensorProduct.ext' fun _ _ => rfl)) + homEquiv_naturality_right := fun _ _ => Action.Hom.ext (ModuleCat.hom_ext (LinearMap.ext + fun _ => LinearMap.ext fun _ => rfl)) })} @[simp] theorem ihom_obj_ρ_def (A B : Rep k G) : ((ihom A).obj B).ρ = ((Rep.ihom A).obj B).ρ := @@ -408,13 +422,13 @@ theorem homEquiv_def (A B C : Rep k G) : (ihom.adjunction A).homEquiv B C = Rep. @[simp] theorem ihom_ev_app_hom (A B : Rep k G) : - Action.Hom.hom ((ihom.ev A).app B) - = TensorProduct.uncurry k A (A →ₗ[k] B) B LinearMap.id.flip := by + Action.Hom.hom ((ihom.ev A).app B) = ModuleCat.ofHom + (TensorProduct.uncurry k A (A →ₗ[k] B) B LinearMap.id.flip) := by ext; rfl @[simp] theorem ihom_coev_app_hom (A B : Rep k G) : - Action.Hom.hom ((ihom.coev A).app B) = (TensorProduct.mk k _ _).flip := - LinearMap.ext fun _ => LinearMap.ext fun _ => rfl + Action.Hom.hom ((ihom.coev A).app B) = ModuleCat.ofHom (TensorProduct.mk k _ _).flip := + ModuleCat.hom_ext <| LinearMap.ext fun _ => LinearMap.ext fun _ => rfl variable (A B C) @@ -435,25 +449,27 @@ variable {A B C} -- `simpNF` times out @[simp, nolint simpNF] theorem MonoidalClosed.linearHomEquiv_hom (f : A ⊗ B ⟶ C) : - (MonoidalClosed.linearHomEquiv A B C f).hom = (TensorProduct.curry f.hom).flip := + (MonoidalClosed.linearHomEquiv A B C f).hom = + ModuleCat.ofHom (TensorProduct.curry f.hom.hom).flip := rfl -- `simpNF` times out @[simp, nolint simpNF] theorem MonoidalClosed.linearHomEquivComm_hom (f : A ⊗ B ⟶ C) : - (MonoidalClosed.linearHomEquivComm A B C f).hom = TensorProduct.curry f.hom := + (MonoidalClosed.linearHomEquivComm A B C f).hom = + ModuleCat.ofHom (TensorProduct.curry f.hom.hom) := rfl theorem MonoidalClosed.linearHomEquiv_symm_hom (f : B ⟶ A ⟶[Rep k G] C) : ((MonoidalClosed.linearHomEquiv A B C).symm f).hom = - TensorProduct.uncurry k A B C f.hom.flip := by + ModuleCat.ofHom (TensorProduct.uncurry k A B C f.hom.hom.flip) := by simp [linearHomEquiv] rfl theorem MonoidalClosed.linearHomEquivComm_symm_hom (f : A ⟶ B ⟶[Rep k G] C) : - ((MonoidalClosed.linearHomEquivComm A B C).symm f).hom - = TensorProduct.uncurry k A B C f.hom := - TensorProduct.ext' fun _ _ => rfl + ((MonoidalClosed.linearHomEquivComm A B C).symm f).hom = + ModuleCat.ofHom (TensorProduct.uncurry k A B C f.hom.hom) := + ModuleCat.hom_ext <| TensorProduct.ext' fun _ _ => rfl end MonoidalClosed @@ -512,8 +528,10 @@ theorem to_Module_monoidAlgebra_map_aux {k G : Type*} [CommRing k] [Monoid G] (V /-- Auxiliary definition for `toModuleMonoidAlgebra`. -/ def toModuleMonoidAlgebraMap {V W : Rep k G} (f : V ⟶ W) : ModuleCat.of (MonoidAlgebra k G) V.ρ.asModule ⟶ ModuleCat.of (MonoidAlgebra k G) W.ρ.asModule := - { f.hom with - map_smul' := fun r x => to_Module_monoidAlgebra_map_aux V.V W.V V.ρ W.ρ f.hom f.comm r x } + ModuleCat.ofHom + { f.hom.hom with + map_smul' := fun r x => to_Module_monoidAlgebra_map_aux V.V W.V V.ρ W.ρ f.hom.hom + (fun g => ModuleCat.hom_ext_iff.mp (f.comm g)) r x } /-- Functorially convert a representation of `G` into a module over `MonoidAlgebra k G`. -/ def toModuleMonoidAlgebra : Rep k G ⥤ ModuleCat.{u} (MonoidAlgebra k G) where @@ -524,8 +542,10 @@ def toModuleMonoidAlgebra : Rep k G ⥤ ModuleCat.{u} (MonoidAlgebra k G) where def ofModuleMonoidAlgebra : ModuleCat.{u} (MonoidAlgebra k G) ⥤ Rep k G where obj M := Rep.of (Representation.ofModule M) map f := - { hom := { f with map_smul' := fun r x => f.map_smul (algebraMap k _ r) x } - comm := fun g => by ext; apply f.map_smul } + { hom := ModuleCat.ofHom + { f.hom with + map_smul' := fun r x => f.hom.map_smul (algebraMap k _ r) x } + comm := fun g => by ext; apply f.hom.map_smul } theorem ofModuleMonoidAlgebra_obj_coe (M : ModuleCat.{u} (MonoidAlgebra k G)) : (ofModuleMonoidAlgebra.obj M : Type u) = RestrictScalars k (MonoidAlgebra k G) M := @@ -551,7 +571,7 @@ def unitIsoAddEquiv {V : Rep k G} : V ≃+ (toModuleMonoidAlgebra ⋙ ofModuleMo /-- Auxiliary definition for `equivalenceModuleMonoidAlgebra`. -/ def counitIso (M : ModuleCat.{u} (MonoidAlgebra k G)) : (ofModuleMonoidAlgebra ⋙ toModuleMonoidAlgebra).obj M ≅ M := - LinearEquiv.toModuleIso' + LinearEquiv.toModuleIso { counitIsoAddEquiv with map_smul' := fun r x => by set_option tactic.skipAssignedInstances false in @@ -566,16 +586,13 @@ theorem unit_iso_comm (V : Rep k G) (g : G) (x : V) : unitIsoAddEquiv ((V.ρ g).toFun x) = ((ofModuleMonoidAlgebra.obj (toModuleMonoidAlgebra.obj V)).ρ g).toFun (unitIsoAddEquiv x) := by dsimp [unitIsoAddEquiv, ofModuleMonoidAlgebra, toModuleMonoidAlgebra] -/- Porting note: rest of broken proof was simp only [AddEquiv.apply_eq_iff_eq, AddEquiv.apply_symm_apply, - Representation.asModuleEquiv_symm_map_rho, Representation.ofModule_asModule_act] -/ - rw [Representation.asModuleEquiv_symm_map_rho] - rfl + Representation.asModuleEquiv_symm_map_rho, Representation.ofModule_asModule_act] /-- Auxiliary definition for `equivalenceModuleMonoidAlgebra`. -/ def unitIso (V : Rep k G) : V ≅ (toModuleMonoidAlgebra ⋙ ofModuleMonoidAlgebra).obj V := Action.mkIso - (LinearEquiv.toModuleIso' + (LinearEquiv.toModuleIso { unitIsoAddEquiv with map_smul' := fun r x => by dsimp [unitIsoAddEquiv] diff --git a/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean b/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean index afc24e101812b..bd80f8783d425 100644 --- a/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean +++ b/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean @@ -257,9 +257,9 @@ include hf private lemma firstRow_exact : (firstRow I M f).Exact where zero k _ := match k with - | 0 => (tens_exact I M f hf).linearMap_comp_eq_zero - | 1 => LinearMap.zero_comp _ - | 2 => LinearMap.zero_comp 0 + | 0 => ModuleCat.hom_ext (tens_exact I M f hf).linearMap_comp_eq_zero + | 1 => ModuleCat.hom_ext (LinearMap.zero_comp _) + | 2 => ModuleCat.hom_ext (LinearMap.zero_comp 0) exact k _ := by rw [ShortComplex.moduleCat_exact_iff] match k with @@ -269,9 +269,9 @@ private lemma firstRow_exact : (firstRow I M f).Exact where private lemma secondRow_exact [Fintype ι] [IsNoetherianRing R] : (secondRow I M f).Exact where zero k _ := match k with - | 0 => (adic_exact I M f hf).linearMap_comp_eq_zero - | 1 => LinearMap.zero_comp (map I f) - | 2 => LinearMap.zero_comp 0 + | 0 => ModuleCat.hom_ext (adic_exact I M f hf).linearMap_comp_eq_zero + | 1 => ModuleCat.hom_ext (LinearMap.zero_comp (map I f)) + | 2 => ModuleCat.hom_ext (LinearMap.zero_comp 0) exact k _ := by rw [ShortComplex.moduleCat_exact_iff] match k with @@ -287,8 +287,8 @@ private def firstRowToSecondRow : firstRow I M f ⟶ secondRow I M f := (ModuleCat.ofHom (ofTensorProduct I M)) (ModuleCat.ofHom 0) (ModuleCat.ofHom 0) - (ofTensorProduct_naturality I <| (LinearMap.ker f).subtype).symm - (ofTensorProduct_naturality I f).symm + (ModuleCat.hom_ext (ofTensorProduct_naturality I <| (LinearMap.ker f).subtype).symm) + (ModuleCat.hom_ext (ofTensorProduct_naturality I f).symm) rfl rfl diff --git a/Mathlib/RingTheory/Coalgebra/TensorProduct.lean b/Mathlib/RingTheory/Coalgebra/TensorProduct.lean index 9efad5fea6593..dafc8270e1007 100644 --- a/Mathlib/RingTheory/Coalgebra/TensorProduct.lean +++ b/Mathlib/RingTheory/Coalgebra/TensorProduct.lean @@ -57,8 +57,8 @@ noncomputable instance TensorProduct.instCoalgebra : Coalgebra R (M ⊗[R] N) := map_comp_comul := by rw [CoalgebraCat.ofComonObjCoalgebraStruct_comul] simp [-Mon_.monMonoidalStruct_tensorObj_X, - ModuleCat.MonoidalCategory.instMonoidalCategoryStruct_tensorHom, - ModuleCat.comp_def, ModuleCat.of, ModuleCat.ofHom, + ModuleCat.MonoidalCategory.instMonoidalCategoryStruct_tensorHom_hom, + ModuleCat.hom_comp, ModuleCat.of, ModuleCat.ofHom, ModuleCat.MonoidalCategory.tensorμ_eq_tensorTensorTensorComm] } end diff --git a/Mathlib/RingTheory/Flat/CategoryTheory.lean b/Mathlib/RingTheory/Flat/CategoryTheory.lean index ea7bc816d2f57..e8aaed01f1d4d 100644 --- a/Mathlib/RingTheory/Flat/CategoryTheory.lean +++ b/Mathlib/RingTheory/Flat/CategoryTheory.lean @@ -53,7 +53,7 @@ lemma iff_lTensor_preserves_shortComplex_exact : fun _ _ _ _ _ _ _ _ _ f g h ↦ moduleCat_exact_iff_function_exact _ |>.1 <| H (.mk (ModuleCat.ofHom f) (ModuleCat.ofHom g) - (DFunLike.ext _ _ h.apply_apply_eq_zero)) + (ModuleCat.hom_ext (DFunLike.ext _ _ h.apply_apply_eq_zero))) (moduleCat_exact_iff_function_exact _ |>.2 h)⟩ lemma iff_rTensor_preserves_shortComplex_exact : @@ -63,7 +63,7 @@ lemma iff_rTensor_preserves_shortComplex_exact : fun _ _ _ _ _ _ _ _ _ f g h ↦ moduleCat_exact_iff_function_exact _ |>.1 <| H (.mk (ModuleCat.ofHom f) (ModuleCat.ofHom g) - (DFunLike.ext _ _ h.apply_apply_eq_zero)) + (ModuleCat.hom_ext (DFunLike.ext _ _ h.apply_apply_eq_zero))) (moduleCat_exact_iff_function_exact _ |>.2 h)⟩ end Module.Flat diff --git a/Mathlib/Topology/Category/Profinite/Nobeling.lean b/Mathlib/Topology/Category/Profinite/Nobeling.lean index c9673426da74d..686efbf1e5992 100644 --- a/Mathlib/Topology/Category/Profinite/Nobeling.lean +++ b/Mathlib/Topology/Category/Profinite/Nobeling.lean @@ -1360,7 +1360,7 @@ theorem succ_mono : CategoryTheory.Mono (ModuleCat.ofHom (πs C o)) := by include hC in theorem succ_exact : (ShortComplex.mk (ModuleCat.ofHom (πs C o)) (ModuleCat.ofHom (Linear_CC' C hsC ho)) - (by ext; apply CC_comp_zero)).Exact := by + (by ext : 2; apply CC_comp_zero)).Exact := by rw [ShortComplex.moduleCat_exact_iff] intro f exact CC_exact C hC hsC ho @@ -1661,7 +1661,7 @@ theorem maxTail_isGood (l : MaxProducts C ho) apply Submodule.add_mem · apply Submodule.finsupp_sum_mem intro q _ - erw [LinearMap.map_smul (fₗ := πs C o) (c := w q) (x := eval (π C (ord I · < o)) q)] + rw [LinearMap.map_smul] apply Submodule.smul_mem apply Submodule.subset_span dsimp only [eval] @@ -1702,7 +1702,7 @@ theorem linearIndependent_comp_of_eval LinearIndependent ℤ (eval (C' C ho)) → LinearIndependent ℤ (ModuleCat.ofHom (Linear_CC' C hsC ho) ∘ SumEval C ho ∘ Sum.inr) := by dsimp [SumEval, ModuleCat.ofHom] - erw [max_eq_eval_unapply C hsC ho] + rw [max_eq_eval_unapply C hsC ho] intro h let f := MaxToGood C hC hsC ho h₁ have hf : f.Injective := maxToGood_injective C hC hsC ho h₁ diff --git a/MathlibTest/CategoryTheory/ConcreteCategory/ModuleCat.lean b/MathlibTest/CategoryTheory/ConcreteCategory/ModuleCat.lean new file mode 100644 index 0000000000000..650d59aeef1ba --- /dev/null +++ b/MathlibTest/CategoryTheory/ConcreteCategory/ModuleCat.lean @@ -0,0 +1,53 @@ +import Mathlib.Algebra.Category.ModuleCat.Basic + +universe v u + +open CategoryTheory ModuleCat + +set_option maxHeartbeats 10000 +set_option synthInstance.maxHeartbeats 2000 + +variable (R : Type u) [CommRing R] + +/- We test if all the coercions and `map_add` lemmas trigger correctly. -/ + +example (X : Type v) [AddCommGroup X] [Module R X] : ⇑(𝟙 (of R X)) = id := by simp + +example {X Y : Type v} [AddCommGroup X] [Module R X] [AddCommGroup Y] [Module R Y] (f : X →ₗ[R] Y) : + ⇑(ModuleCat.ofHom f) = ⇑f := by simp + +example {X Y : Type v} [AddCommGroup X] [Module R X] [AddCommGroup Y] [Module R Y] (f : X →ₗ[R] Y) + (x : X) : (ModuleCat.ofHom f) x = f x := by simp + +example {X Y Z : ModuleCat R} (f : X ⟶ Y) (g : Y ⟶ Z) : ⇑(f ≫ g) = ⇑g ∘ ⇑f := by simp + +example {X Y Z : Type v} [AddCommGroup X] [Module R X] [AddCommGroup Y] [Module R Y] [Ring Z] + [Algebra R Z] (f : X →ₗ[R] Y) (g : Y →ₗ[R] Z) : + ⇑(ModuleCat.ofHom f ≫ ModuleCat.ofHom g) = g ∘ f := by simp + +example {X Y : Type v} [AddCommGroup X] [Module R X] [AddCommGroup Y] [Module R Y] {Z : ModuleCat R} + (f : X →ₗ[R] Y) (g : of R Y ⟶ Z) : + ⇑(ModuleCat.ofHom f ≫ g) = g ∘ f := by simp + +example {X Y : ModuleCat R} {Z : Type v} [Ring Z] [Algebra R Z] (f : X ⟶ Y) (g : Y ⟶ of R Z) : + ⇑(f ≫ g) = g ∘ f := by simp + +example {Y Z : ModuleCat R} {X : Type v} [AddCommGroup X] [Module R X] (f : of R X ⟶ Y) (g : Y ⟶ Z) : + ⇑(f ≫ g) = g ∘ f := by simp + +example {X Y Z : ModuleCat R} (f : X ⟶ Y) (g : Y ⟶ Z) (x : X) : (f ≫ g) x = g (f x) := by simp + +example {X Y : ModuleCat R} (e : X ≅ Y) (x : X) : e.inv (e.hom x) = x := by simp + +example {X Y : ModuleCat R} (e : X ≅ Y) (y : Y) : e.hom (e.inv y) = y := by simp + +example (X : ModuleCat R) : ⇑(𝟙 X) = id := by simp + +example {M N : ModuleCat.{v} R} (f : M ⟶ N) (x y : M) : f (x + y) = f x + f y := by + simp + +example {M N : ModuleCat.{v} R} (f : M ⟶ N) : f 0 = 0 := by + simp + +example {M N : ModuleCat.{v} R} (f : M ⟶ N) (r : R) (m : M) : f (r • m) = r • f m := by + simp From 67d092e380faf850fd6a1febf81c27fb33c6c5c8 Mon Sep 17 00:00:00 2001 From: Scott Carnahan <128885296+ScottCarnahan@users.noreply.github.com> Date: Thu, 5 Dec 2024 10:46:42 +0000 Subject: [PATCH 537/829] feat(Algebra/Vertex): Define vertex operators (#19298) A vertex operator for an `R`-module `V` is an `R`-module map from `V` to Laurent series on `V`. This PR defines vertex operators and introduces a normalized indexing for coefficients. --- Mathlib.lean | 1 + Mathlib/Algebra/Vertex/HVertexOperator.lean | 10 +++ Mathlib/Algebra/Vertex/VertexOperator.lean | 99 +++++++++++++++++++++ docs/references.bib | 21 +++++ 4 files changed, 131 insertions(+) create mode 100644 Mathlib/Algebra/Vertex/VertexOperator.lean diff --git a/Mathlib.lean b/Mathlib.lean index f71ab46717573..698f1cd681d1f 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -938,6 +938,7 @@ import Mathlib.Algebra.Tropical.Basic import Mathlib.Algebra.Tropical.BigOperators import Mathlib.Algebra.Tropical.Lattice import Mathlib.Algebra.Vertex.HVertexOperator +import Mathlib.Algebra.Vertex.VertexOperator import Mathlib.AlgebraicGeometry.AffineScheme import Mathlib.AlgebraicGeometry.AffineSpace import Mathlib.AlgebraicGeometry.Cover.MorphismProperty diff --git a/Mathlib/Algebra/Vertex/HVertexOperator.lean b/Mathlib/Algebra/Vertex/HVertexOperator.lean index 4248cbf1e971d..d68241ff4baf6 100644 --- a/Mathlib/Algebra/Vertex/HVertexOperator.lean +++ b/Mathlib/Algebra/Vertex/HVertexOperator.lean @@ -92,6 +92,16 @@ def of_coeff (f : Γ → V →ₗ[R] W) @[deprecated (since := "2024-06-18")] alias _root_.VertexAlg.HetVertexOperator.of_coeff := of_coeff +@[simp] +theorem add_coeff (A B : HVertexOperator Γ R V W) : (A + B).coeff = A.coeff + B.coeff := by + ext + simp + +@[simp] +theorem smul_coeff (A : HVertexOperator Γ R V W) (r : R) : (r • A).coeff = r • (A.coeff) := by + ext + simp + end Coeff section Products diff --git a/Mathlib/Algebra/Vertex/VertexOperator.lean b/Mathlib/Algebra/Vertex/VertexOperator.lean new file mode 100644 index 0000000000000..e7eaef0b16c08 --- /dev/null +++ b/Mathlib/Algebra/Vertex/VertexOperator.lean @@ -0,0 +1,99 @@ +/- +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.Vertex.HVertexOperator +import Mathlib.Data.Int.Interval + +/-! +# Vertex operators +In this file we introduce vertex operators as linear maps to Laurent series. + +## Definitions + * `VertexOperator` is an `R`-linear map from an `R`-module `V` to `LaurentSeries V`. + * `VertexOperator.ncoeff` is the coefficient of a vertex operator under normalized indexing. + +## TODO + * `HasseDerivative` : A divided-power derivative. + * `Locality` : A weak form of commutativity. + * `Residue products` : A family of products on `VertexOperator R V` parametrized by integers. + +## References + * [G. Mason, *Vertex rings and Pierce bundles*][mason2017] + * [A. Matsuo, K. Nagatomo, *On axioms for a vertex algebra and locality of quantum + fields*][matsuo1997] +-/ + +noncomputable section + +variable {R V : Type*} [CommRing R] [AddCommGroup V] [Module R V] + +/-- A vertex operator over a commutative ring `R` is an `R`-linear map from an `R`-module `V` to +Laurent series with coefficients in `V`. We write this as a specialization of the heterogeneous +case. -/ +abbrev VertexOperator (R : Type*) (V : Type*) [CommRing R] [AddCommGroup V] + [Module R V] := HVertexOperator ℤ R V V + +namespace VertexOperator + +open HVertexOperator + +@[ext] +theorem ext (A B : VertexOperator R V) (h : ∀ v : V, A v = B v) : + A = B := LinearMap.ext h + +/-- The coefficient of a vertex operator under normalized indexing. -/ +def ncoeff {R} [CommRing R] [AddCommGroup V] [Module R V] (A : VertexOperator R V) (n : ℤ) : + Module.End R V := HVertexOperator.coeff A (-n - 1) + +/-- In the literature, the `n`th normalized coefficient of a vertex operator `A` is written as +either `Aₙ` or `A(n)`. -/ +scoped[VertexOperator] notation A "[[" n "]]" => ncoeff A n + +@[simp] +theorem coeff_eq_ncoeff (A : VertexOperator R V) + (n : ℤ) : HVertexOperator.coeff A n = A [[-n - 1]] := by + rw [ncoeff, neg_sub, sub_neg_eq_add, add_sub_cancel_left] + +@[simp] +theorem ncoeff_add (A B : VertexOperator R V) (n : ℤ) : (A + B) [[n]] = A [[n]] + B [[n]] := by + rw [ncoeff, ncoeff, ncoeff, add_coeff, Pi.add_apply] + +@[simp] +theorem ncoeff_smul (A : VertexOperator R V) (r : R) (n : ℤ) : (r • A) [[n]] = r • A [[n]] := by + rw [ncoeff, ncoeff, smul_coeff, Pi.smul_apply] + +theorem ncoeff_eq_zero_of_lt_order (A : VertexOperator R V) (n : ℤ) (x : V) + (h : -n - 1 < HahnSeries.order ((HahnModule.of R).symm (A x))) : (A [[n]]) x = 0 := by + simp only [ncoeff, HVertexOperator.coeff, LinearMap.coe_mk, AddHom.coe_mk] + exact HahnSeries.coeff_eq_zero_of_lt_order h + +theorem coeff_eq_zero_of_lt_order (A : VertexOperator R V) (n : ℤ) (x : V) + (h : n < HahnSeries.order ((HahnModule.of R).symm (A x))) : coeff A n x = 0 := by + rw [coeff_eq_ncoeff, ncoeff_eq_zero_of_lt_order A (-n - 1) x] + omega + +/-- Given an endomorphism-valued function on integers satisfying a pointwise bounded-pole condition, +we produce a vertex operator. -/ +noncomputable def of_coeff (f : ℤ → Module.End R V) + (hf : ∀ (x : V), ∃ (n : ℤ), ∀ (m : ℤ), m < n → (f m) x = 0) : VertexOperator R V := + HVertexOperator.of_coeff f + (fun x => HahnSeries.suppBddBelow_supp_PWO (fun n => (f n) x) + (HahnSeries.forallLTEqZero_supp_BddBelow (fun n => (f n) x) + (Exists.choose (hf x)) (Exists.choose_spec (hf x)))) + +@[simp] +theorem of_coeff_apply_coeff (f : ℤ → Module.End R V) + (hf : ∀ (x : V), ∃ n, ∀ m < n, (f m) x = 0) (x : V) (n : ℤ) : + ((HahnModule.of R).symm ((of_coeff f hf) x)).coeff n = (f n) x := by + rfl + +@[simp] +theorem ncoeff_of_coeff (f : ℤ → Module.End R V) + (hf : ∀(x : V), ∃(n : ℤ), ∀(m : ℤ), m < n → (f m) x = 0) (n : ℤ) : + (of_coeff f hf) [[n]] = f (-n - 1) := by + ext v + rw [ncoeff, coeff_apply, of_coeff_apply_coeff] + +end VertexOperator diff --git a/docs/references.bib b/docs/references.bib index 976afa9ae6d69..69f11f4a17b95 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -2602,6 +2602,16 @@ @Article{ markowsky1976 pages = {53-68} } +@Article{ mason2017, + title = {Vertex rings and Pierce bundles}, + author = {Geoffrey Mason}, + year = {2017}, + eprint = {1707.00328}, + archiveprefix = {arXiv}, + primaryclass = {math.RA}, + url = {https://arxiv.org/abs/1707.00328v1} +} + @InProceedings{ mathlib2020, author = {The mathlib Community}, title = {The Lean Mathematical Library}, @@ -2620,6 +2630,17 @@ @InProceedings{ mathlib2020 series = {CPP 2020} } +@Article{ matsuo1997, + title = {On axioms for a vertex algebra and locality of quantum + fields}, + author = {Atsushi Matsuo, Kiyokazu Nagatomo}, + year = {1997}, + eprint = {9706118}, + archiveprefix = {arXiv}, + primaryclass = {hep-th}, + url = {https://arxiv.org/abs/hep-th/9706118} +} + @Book{ mattila1995, author = {Mattila, Pertti}, title = {Geometry of sets and measures in {E}uclidean spaces}, From 0331c85a6e52d2f2241d9ee2d02808a97b70708a Mon Sep 17 00:00:00 2001 From: Gabin Date: Thu, 5 Dec 2024 16:32:07 +0000 Subject: [PATCH 538/829] chore (ModelTheory/Fraisse): simplify IsFraisse definition (#19041) Remove the condition of being equiv_invariant in IsFraisse Deletions: - IsFraisse.is_equiv_invariant Co-authored-by: Gabin <68381468+Vegan-Parabola@users.noreply.github.com> --- Mathlib/ModelTheory/Fraisse.lean | 22 ++++++++++++---------- 1 file changed, 12 insertions(+), 10 deletions(-) diff --git a/Mathlib/ModelTheory/Fraisse.lean b/Mathlib/ModelTheory/Fraisse.lean index ce7356ee4aeb5..11b8204c05fa6 100644 --- a/Mathlib/ModelTheory/Fraisse.lean +++ b/Mathlib/ModelTheory/Fraisse.lean @@ -29,8 +29,8 @@ Fraïssé limit - the countable ultrahomogeneous structure with that age. - A class `K` has `FirstOrder.Language.Amalgamation` when for any pair of embeddings of a structure `M` in `K` into other structures in `K`, those two structures can be embedded into a fourth structure in `K` such that the resulting square of embeddings commutes. -- `FirstOrder.Language.IsFraisse` indicates that a class is nonempty, isomorphism-invariant, - essentially countable, and satisfies the hereditary, joint embedding, and amalgamation properties. +- `FirstOrder.Language.IsFraisse` indicates that a class is nonempty, essentially countable, + and satisfies the hereditary, joint embedding, and amalgamation properties. - `FirstOrder.Language.IsFraisseLimit` indicates that a structure is a Fraïssé limit for a given class. @@ -109,12 +109,11 @@ def Amalgamation : Prop := M ∈ K → N ∈ K → P ∈ K → ∃ (Q : Bundled.{w} L.Structure) (NQ : N ↪[L] Q) (PQ : P ↪[L] Q), Q ∈ K ∧ NQ.comp MN = PQ.comp MP -/-- A Fraïssé class is a nonempty, isomorphism-invariant, essentially countable class of structures -satisfying the hereditary, joint embedding, and amalgamation properties. -/ +/-- A Fraïssé class is a nonempty, essentially countable class of structures satisfying the +hereditary, joint embedding, and amalgamation properties. -/ class IsFraisse : Prop where is_nonempty : K.Nonempty FG : ∀ M : Bundled.{w} L.Structure, M ∈ K → Structure.FG L M - is_equiv_invariant : ∀ M N : Bundled.{w} L.Structure, Nonempty (M ≃[L] N) → (M ∈ K ↔ N ∈ K) is_essentially_countable : (Quotient.mk' '' K).Countable hereditary : Hereditary K jointEmbedding : JointEmbedding K @@ -146,6 +145,10 @@ theorem Hereditary.is_equiv_invariant_of_fg (h : Hereditary K) ⟨fun MK => h M MK ((fg M MK).mem_age_of_equiv hn), fun NK => h N NK ((fg N NK).mem_age_of_equiv ⟨hn.some.symm⟩)⟩ +theorem IsFraisse.is_equiv_invariant [h : IsFraisse K] {M N : Bundled.{w} L.Structure} + (hn : Nonempty (M ≃[L] N)) : M ∈ K ↔ N ∈ K := + h.hereditary.is_equiv_invariant_of_fg h.FG M N hn + variable (M) theorem age.nonempty : (L.age M).Nonempty := @@ -221,7 +224,6 @@ theorem age_directLimit {ι : Type w} [Preorder ι] [IsDirected ι (· ≤ ·)] /-- Sufficient conditions for a class to be the age of a countably-generated structure. -/ theorem exists_cg_is_age_of (hn : K.Nonempty) - (h : ∀ M N : Bundled.{w} L.Structure, Nonempty (M ≃[L] N) → (M ∈ K ↔ N ∈ K)) (hc : (Quotient.mk' '' K).Countable) (fg : ∀ M : Bundled.{w} L.Structure, M ∈ K → Structure.FG L M) (hp : Hereditary K) (jep : JointEmbedding K) : ∃ M : Bundled.{w} L.Structure, Structure.CG L M ∧ L.age M = K := by @@ -234,7 +236,7 @@ theorem exists_cg_is_age_of (hn : K.Nonempty) -- Porting note: fix hP2 because `Quotient.out (Quotient.mk' x) ≈ a` was not simplified -- to `x ≈ a` in hF replace hP2 := Setoid.trans (Setoid.symm (Quotient.mk_out P)) hP2 - exact (h _ _ hP2).1 hP1 + exact (hp.is_equiv_invariant_of_fg fg _ _ hP2).1 hP1 choose P hPK hP hFP using fun (N : K) (n : ℕ) => jep N N.2 (F (n + 1)).out (hF' _) let G : ℕ → K := @Nat.rec (fun _ => K) ⟨(F 0).out, hF' 0⟩ fun n N => ⟨P N n, hPK N n⟩ -- Poting note: was @@ -265,8 +267,8 @@ theorem exists_countable_is_age_of_iff [Countable (Σ l, L.Functions l)] : · rintro ⟨M, h1, h2, rfl⟩ refine ⟨age.nonempty M, age.is_equiv_invariant L M, age.countable_quotient M, fun N hN => hN.1, age.hereditary M, age.jointEmbedding M⟩ - · rintro ⟨Kn, eqinv, cq, hfg, hp, jep⟩ - obtain ⟨M, hM, rfl⟩ := exists_cg_is_age_of Kn eqinv cq hfg hp jep + · rintro ⟨Kn, _, cq, hfg, hp, jep⟩ + obtain ⟨M, hM, rfl⟩ := exists_cg_is_age_of Kn cq hfg hp jep exact ⟨M, Structure.cg_iff_countable.1 hM, rfl⟩ variable (L) @@ -358,7 +360,7 @@ theorem IsUltrahomogeneous.amalgamation_age (h : L.IsUltrahomogeneous M) : theorem IsUltrahomogeneous.age_isFraisse [Countable M] (h : L.IsUltrahomogeneous M) : IsFraisse (L.age M) := - ⟨age.nonempty M, fun _ hN => hN.1, age.is_equiv_invariant L M, age.countable_quotient M, + ⟨age.nonempty M, fun _ hN => hN.1, age.countable_quotient M, age.hereditary M, age.jointEmbedding M, h.amalgamation_age⟩ namespace IsFraisseLimit From 36c7c9d4fa8fe9102b8304fccfb4b45eb164e28a Mon Sep 17 00:00:00 2001 From: Christopher Hoskin Date: Thu, 5 Dec 2024 16:41:15 +0000 Subject: [PATCH 539/829] feat(Algebra/BigOperators/Group/Finset) : Finset.prod_disjoint_filters (#19712) Decomposes the product(sum) over two filters where they are pointwise disjoint. Used in #19432 Co-authored-by: Christopher Hoskin --- Mathlib/Algebra/BigOperators/Group/Finset.lean | 7 +++++++ Mathlib/Data/Finset/Filter.lean | 9 +++++++++ 2 files changed, 16 insertions(+) diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index c235e6c2bd74a..34cd0b147e819 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -426,6 +426,13 @@ lemma prod_filter_not_mul_prod_filter (s : Finset α) (p : α → Prop) [Decidab (∏ x ∈ s.filter fun x ↦ ¬p x, f x) * ∏ x ∈ s.filter p, f x = ∏ x ∈ s, f x := by rw [mul_comm, prod_filter_mul_prod_filter_not] +@[to_additive] +theorem prod_filter_xor (p q : α → Prop) [DecidableEq α] [DecidablePred p] [DecidablePred q] : + (∏ x ∈ s with (Xor' (p x) (q x)), f x) = + (∏ x ∈ s with (p x ∧ ¬ q x), f x) * (∏ x ∈ s with (q x ∧ ¬ p x), f x) := by + rw [← prod_union (disjoint_filter_and_not_filter _ _), ← filter_or] + rfl + section ToList @[to_additive (attr := simp)] diff --git a/Mathlib/Data/Finset/Filter.lean b/Mathlib/Data/Finset/Filter.lean index d42d8d22414fa..746efb1723405 100644 --- a/Mathlib/Data/Finset/Filter.lean +++ b/Mathlib/Data/Finset/Filter.lean @@ -214,6 +214,15 @@ lemma _root_.Set.pairwiseDisjoint_filter [DecidableEq β] (f : α → β) (s : S obtain ⟨-, rfl⟩ : x ∈ t ∧ f x = j := by simpa using hj hx contradiction +theorem disjoint_filter_and_not_filter : + Disjoint (s.filter (fun x ↦ p x ∧ ¬q x)) (s.filter (fun x ↦ q x ∧ ¬p x)) := by + intro _ htp htq + simp [bot_eq_empty, le_eq_subset, subset_empty] + by_contra hn + rw [← not_nonempty_iff_eq_empty, not_not] at hn + obtain ⟨_, hx⟩ := hn + exact (mem_filter.mp (htq hx)).2.2 (mem_filter.mp (htp hx)).2.1 + variable {p q} lemma filter_inj : s.filter p = t.filter p ↔ ∀ ⦃a⦄, p a → (a ∈ s ↔ a ∈ t) := by From d070358a60c92d366214c49fa4731575a4c9d651 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 5 Dec 2024 22:47:52 +0000 Subject: [PATCH 540/829] chore: run lean4checker on most recent nightly-testing-YYYY-MM-DD tag (#19735) Hopefully this would have given us an earlier heads-up about the recent lean4checker problem. --- .github/workflows/lean4checker.yml | 49 ++++++++++++++---------------- 1 file changed, 23 insertions(+), 26 deletions(-) diff --git a/.github/workflows/lean4checker.yml b/.github/workflows/lean4checker.yml index 066648826282a..c521c9164c994 100644 --- a/.github/workflows/lean4checker.yml +++ b/.github/workflows/lean4checker.yml @@ -1,4 +1,3 @@ -# https://chat.openai.com/share/26102e95-ac9c-4d03-a000-73e4f6cee8cd name: lean4checker Workflow on: @@ -6,14 +5,19 @@ on: - cron: '0 0 * * *' # Runs at 00:00 UTC every day workflow_dispatch: +env: + DEFAULT_BRANCH: master + TAG_PATTERN: '^nightly-testing-[0-9]{4}-[0-9]{2}-[0-9]{2}$' + jobs: check-lean4checker: - runs-on: pr + runs-on: ubuntu-latest + strategy: + matrix: + branch_type: [master, nightly] steps: - - - name: cleanup + - name: Cleanup run: | - find . -name . -o -prune -exec rm -rf -- {} + # Delete all but the 5 most recent toolchains. # Make sure to delete both the `~/.elan/toolchains/X` directory and the `~/.elan/update-hashes/X` file. # Skip symbolic links (`-type d`), the current directory (`! -name .`), and `nightly` and `stable`. @@ -23,21 +27,24 @@ jobs: : # Do nothing on failure, but suppress errors fi - # The Hoskinson runners may not have jq installed, so do that now. - - name: 'Setup jq' - uses: dcarbone/install-jq-action@v2.1.0 - - - name: install elan + - name: Install elan run: | set -o pipefail curl -sSfL https://github.com/leanprover/elan/releases/download/v3.1.1/elan-x86_64-unknown-linux-gnu.tar.gz | tar xz ./elan-init -y --default-toolchain none echo "$HOME/.elan/bin" >> "${GITHUB_PATH}" - - name: Checkout master branch + - name: Fetch latest tags (if nightly) + if: matrix.branch_type == 'nightly' + run: | + git fetch --tags + LATEST_TAG=$(git tag | grep -E "${{ env.TAG_PATTERN }}" | sort -r | head -n 1) + echo "LATEST_TAG=${LATEST_TAG}" >> "$GITHUB_ENV" + + - name: Checkout branch or tag uses: actions/checkout@v4 with: - ref: 'master' + ref: ${{ matrix.branch_type == 'master' && env.DEFAULT_BRANCH || env.LATEST_TAG }} - name: If using a lean-pr-release toolchain, uninstall run: | @@ -46,7 +53,7 @@ jobs: elan toolchain uninstall "$(cat lean-toolchain)" fi - - name: print lean and lake versions + - name: Print Lean and Lake versions run: | lean --version lake --version @@ -55,16 +62,6 @@ jobs: run: | lake exe cache get - - name: prune ProofWidgets .lake - run: | - # The ProofWidgets release contains not just the `.js` (which we need in order to build) - # but also `.oleans`, which may have been built with the wrong toolchain. - # This removes them. - # See discussion at https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/nightly-testing/near/411225235 - rm -rf .lake/packages/proofwidgets/.lake/build/lib - rm -rf .lake/packages/proofwidgets/.lake/build/ir - lake build ProofWidgets - - name: Check environments using lean4checker id: lean4checker run: | @@ -100,7 +97,7 @@ jobs: type: 'stream' topic: 'lean4checker' content: | - ✅ lean4checker [succeeded](https://github.com/${{ github.repository }}/actions/runs/${{ github.run_id }}) on ${{ github.sha }} + ✅ lean4checker [succeeded](https://github.com/${{ github.repository }}/actions/runs/${{ github.run_id }}) on ${{ github.sha }} (branch: ${{ matrix.branch_type == 'master' && 'master' || env.LATEST_TAG }}) - name: Post failure message on Zulip if: failure() @@ -113,7 +110,7 @@ jobs: type: 'stream' topic: 'lean4checker failure' content: | - ❌ lean4checker [failed](https://github.com/${{ github.repository }}/actions/runs/${{ github.run_id }}) on ${{ github.sha }} + ❌ lean4checker [failed](https://github.com/${{ github.repository }}/actions/runs/${{ github.run_id }}) on ${{ github.sha }} (branch: ${{ matrix.branch_type == 'master' && 'master' || env.LATEST_TAG }}) continue-on-error: true - name: Post failure message on Zulip main topic @@ -127,5 +124,5 @@ jobs: type: 'stream' topic: 'lean4checker' content: | - ❌ lean4checker [failed](https://github.com/${{ github.repository }}/actions/runs/${{ github.run_id }}) on ${{ github.sha }} + ❌ lean4checker [failed](https://github.com/${{ github.repository }}/actions/runs/${{ github.run_id }}) on ${{ github.sha }} (branch: ${{ matrix.branch_type == 'master' && 'master' || env.LATEST_TAG }}) continue-on-error: true From 0ae2683abccf50022a5c3c5e48a1c51899e91647 Mon Sep 17 00:00:00 2001 From: Mitchell Lee Date: Fri, 6 Dec 2024 05:31:08 +0000 Subject: [PATCH 541/829] feat: power of product of two reflections (#13270) Let $M$ be a module over a commutative ring $R$. Let $x, y \in M$ and $f, g \in M^*$ with $f(x) = g(y) = 2$. The corresponding reflections $r_1, r_2 \colon M \to M$ are given by $r_1z = z - f(z) x$ and $r_2 z = z - g(z) y$. These are linear automorphisms of $M$. We prove that for all $z \in M$ and all $m \in \mathbb{Z}$, we have: $$(r_1 r_2)^m z = z + S_{\left\lfloor \frac{m-2}{2} \right \rfloor}(t) (S_{\left\lfloor \frac{m-1}{2} \right \rfloor}(t) + S_{\left\lfloor \frac{m-3}{2} \right \rfloor}(t)) (g(x) f(z) y - g(z) y - f(z) x) + S_{\left\lfloor \frac{m-1}{2} \right \rfloor}(t) (S_{\left\lfloor \frac{m}{2} \right \rfloor}(t) + S_{\left\lfloor \frac{m-2}{2} \right \rfloor}(t)) (f(y) g(z) x - f(z) x - g(z) y)$$ $$(r_1 r_2)^m x = (S_{m}(t) + S_{m - 1}(t)) x - S_{m-1} g(x) y$$ $$r_2(r_1 r_2)^m x = (S_{m}(t) + S_{m - 1}(t)) x - S_{m} g(x) y$$ where $t = f(y) g(x) - 2$ and $S_n$ refers to a Chebyshev $S$-polynomial. These formulas may look ridiculous, but they are necessary for constructing reflection representations of a Coxeter group over an arbitrary ring. See #13291 for more details. Co-authored-by: Johan Commelin Co-authored-by: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> --- Mathlib/Algebra/Module/Equiv/Basic.lean | 2 + Mathlib/LinearAlgebra/Reflection.lean | 197 ++++++++++++++++++- Mathlib/RingTheory/Polynomial/Chebyshev.lean | 10 + 3 files changed, 208 insertions(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Module/Equiv/Basic.lean b/Mathlib/Algebra/Module/Equiv/Basic.lean index 67ae9340c5ef4..1ccd71d6cbcbb 100644 --- a/Mathlib/Algebra/Module/Equiv/Basic.lean +++ b/Mathlib/Algebra/Module/Equiv/Basic.lean @@ -96,6 +96,8 @@ theorem coe_pow (e : M ≃ₗ[R] M) (n : ℕ) : ⇑(e ^ n) = e^[n] := hom_coe_po theorem pow_apply (e : M ≃ₗ[R] M) (n : ℕ) (m : M) : (e ^ n) m = e^[n] m := congr_fun (coe_pow e n) m +@[simp] lemma mul_apply (f : M ≃ₗ[R] M) (g : M ≃ₗ[R] M) (x : M) : (f * g) x = f (g x) := rfl + /-- Restriction from `R`-linear automorphisms of `M` to `R`-linear endomorphisms of `M`, promoted to a monoid hom. -/ @[simps] diff --git a/Mathlib/LinearAlgebra/Reflection.lean b/Mathlib/LinearAlgebra/Reflection.lean index d807bcfeec42e..79d13d900e5b4 100644 --- a/Mathlib/LinearAlgebra/Reflection.lean +++ b/Mathlib/LinearAlgebra/Reflection.lean @@ -1,12 +1,13 @@ /- Copyright (c) 2023 Oliver Nash. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Oliver Nash, Deepro Choudhury +Authors: Oliver Nash, Deepro Choudhury, Mitchell Lee, Johan Commelin -/ import Mathlib.Algebra.Module.LinearMap.Basic import Mathlib.GroupTheory.OrderOfElement import Mathlib.LinearAlgebra.Dual import Mathlib.LinearAlgebra.FiniteSpan +import Mathlib.RingTheory.Polynomial.Chebyshev /-! # Reflections in linear algebra @@ -28,6 +29,9 @@ is characterised by properties 1 and 2 above, and is a linear isometry. * `Module.reflection`: the definition of the map `y ↦ y - (f y) • x`. This requires the assumption that `f x = 2` but by way of compensation it produces a linear equivalence rather than a mere linear map. + * `Module.reflection_mul_reflection_pow_apply`: a formula for $(r_1 r_2)^m z$, where $r_1$ and + $r_2$ are reflections and $z \in M$. It involves the Chebyshev polynomials and holds over any + commutative ring. This is used to define reflection representations of Coxeter groups. * `Module.Dual.eq_of_preReflection_mapsTo`: a uniqueness result about reflections preserving finite spanning sets that is useful in the theory of root data / systems. @@ -101,6 +105,9 @@ lemma involutive_reflection (h : f x = 2) : Involutive (reflection h) := involutive_preReflection h +@[simp] +lemma reflection_inv (h : f x = 2) : (reflection h)⁻¹ = reflection h := rfl + @[simp] lemma reflection_symm (h : f x = 2) : (reflection h).symm = reflection h := @@ -114,6 +121,194 @@ lemma bijOn_reflection_of_mapsTo {Φ : Set M} (h : f x = 2) (h' : MapsTo (reflec BijOn (reflection h) Φ Φ := (invOn_reflection_of_mapsTo h).bijOn h' h' +/-! ### Powers of the product of two reflections + +Let $M$ be a module over a commutative ring $R$. Let $x, y \in M$ and $f, g \in M^*$ with +$f(x) = g(y) = 2$. The corresponding reflections $r_1, r_2 \colon M \to M$ (`Module.reflection`) are +given by $r_1z = z - f(z) x$ and $r_2 z = z - g(z) y$. These are linear automorphisms of $M$. + +To define reflection representations of a Coxeter group, it is important to be able to compute the +order of the composition $r_1 r_2$. + +Note that if $M$ is a real inner product space and $r_1$ and $r_2$ are both orthogonal +reflections (i.e. $f(z) = 2 \langle x, z \rangle / \langle x, x \rangle$ and +$g(z) = 2 \langle y, z\rangle / \langle y, y\rangle$ for all $z \in M$), +then $r_1 r_2$ is a rotation by the angle +$$\cos^{-1}\left(\frac{f(y) g(x) - 2}{2}\right)$$ +and one may determine the order of $r_1 r_2$ accordingly. + +However, if $M$ does not have an inner product, and even if $R$ is not $\mathbb{R}$, then we may +instead use the formulas in this section. These formulas all involve evaluating Chebyshev +$S$-polynomials (`Polynomial.Chebyshev.S`) at $t = f(y) g(x) - 2$, and they hold over any +commutative ring. -/ +section + +open Int Polynomial.Chebyshev + +variable {x y : M} {f g : Dual R M} (hf : f x = 2) (hg : g y = 2) + +/-- A formula for $(r_1 r_2)^m z$, where $m$ is a natural number and $z \in M$. -/ +lemma reflection_mul_reflection_pow_apply (m : ℕ) (z : M) + (t : R := f y * g x - 2) (ht : t = f y * g x - 2 := by rfl) : + ((reflection hf * reflection hg) ^ m) z = + z + + ((S R ((m - 2) / 2)).eval t * ((S R ((m - 1) / 2)).eval t + (S R ((m - 3) / 2)).eval t)) • + ((g x * f z - g z) • y - f z • x) + + ((S R ((m - 1) / 2)).eval t * ((S R (m / 2)).eval t + (S R ((m - 2) / 2)).eval t)) • + ((f y * g z - f z) • x - g z • y) := by + induction m with + | zero => simp + | succ m ih => + /- Now, let us collect two facts about the evaluations of `S r k`. These easily follow from the + properties of the `S` polynomials. -/ + have S_eval_t_sub_two (k : ℤ) : + (S R (k - 2)).eval t = t * (S R (k - 1)).eval t - (S R k).eval t := by + simp [S_sub_two] + have S_eval_t_sq_add_S_eval_t_sq (k : ℤ) : + (S R k).eval t ^ 2 + (S R (k + 1)).eval t ^ 2 - t * (S R k).eval t * (S R (k + 1)).eval t + = 1 := by + simpa using congr_arg (Polynomial.eval t) (S_sq_add_S_sq R k) + -- Apply the inductive hypothesis. + rw [pow_succ', LinearEquiv.mul_apply, ih, LinearEquiv.mul_apply] + -- Expand out all the reflections and use `hf`, `hg`. + simp only [reflection_apply, map_add, map_sub, map_smul, hf, hg] + -- `m` can be written in the form `2 * k + e`, where `e` is `0` or `1`. + push_cast + set k : ℤ := m / 2 + set e : ℤ := m % 2 + rw [show m = 2 * k + e from (Int.ediv_add_emod m 2).symm] + simp_rw [add_assoc (2 * k), add_sub_assoc (2 * k), add_comm (2 * k), + add_mul_ediv_left _ k (by norm_num : (2 : ℤ) ≠ 0)] + have he : e = 0 ∨ e = 1 := by omega + clear_value e + /- Now, equate the coefficients on both sides. These linear combinations were + found using `polyrith`. -/ + match_scalars + · rfl + · linear_combination (norm := skip) (-g z * f y * (S R (e - 1 + k)).eval t + + f z * (S R (e - 1 + k)).eval t) * S_eval_t_sub_two (e + k) + + (-g z * f y + f z) * S_eval_t_sq_add_S_eval_t_sq (k - 1) + subst ht + obtain rfl | rfl : e = 0 ∨ e = 1 := he <;> ring_nf + · linear_combination (norm := skip) + g z * (S R (e - 1 + k)).eval t * S_eval_t_sub_two (e + k) + + g z * S_eval_t_sq_add_S_eval_t_sq (k - 1) + subst ht + obtain rfl | rfl : e = 0 ∨ e = 1 := he <;> ring_nf + +/-- A formula for $(r_1 r_2)^m$, where $m$ is a natural number. -/ +lemma reflection_mul_reflection_pow (m : ℕ) + (t : R := f y * g x - 2) (ht : t = f y * g x - 2 := by rfl) : + ((reflection hf * reflection hg) ^ m).toLinearMap = + LinearMap.id (R := R) (M := M) + + ((S R ((m - 2) / 2)).eval t * ((S R ((m - 1) / 2)).eval t + (S R ((m - 3) / 2)).eval t)) • + ((g x • f - g).smulRight y - f.smulRight x) + + ((S R ((m - 1) / 2)).eval t * ((S R (m / 2)).eval t + (S R ((m - 2) / 2)).eval t)) • + ((f y • g - f).smulRight x - g.smulRight y) := by + ext z + simpa using reflection_mul_reflection_pow_apply hf hg m z t ht + +/-- A formula for $(r_1 r_2)^m z$, where $m$ is an integer and $z \in M$. -/ +lemma reflection_mul_reflection_zpow_apply (m : ℤ) (z : M) + (t : R := f y * g x - 2) (ht : t = f y * g x - 2 := by rfl) : + ((reflection hf * reflection hg) ^ m) z = + z + + ((S R ((m - 2) / 2)).eval t * ((S R ((m - 1) / 2)).eval t + (S R ((m - 3) / 2)).eval t)) • + ((g x * f z - g z) • y - f z • x) + + ((S R ((m - 1) / 2)).eval t * ((S R (m / 2)).eval t + (S R ((m - 2) / 2)).eval t)) • + ((f y * g z - f z) • x - g z • y) := by + induction m using Int.negInduction with + | nat m => exact_mod_cast reflection_mul_reflection_pow_apply hf hg m z t ht + | neg _ m => + have ht' : t = g x * f y - 2 := by rwa [mul_comm (g x)] + rw [zpow_neg, ← inv_zpow, mul_inv_rev, reflection_inv, reflection_inv, zpow_natCast, + reflection_mul_reflection_pow_apply hg hf m z t ht', add_right_comm z] + have aux {a b : ℤ} (hab : a + b = -3) : a / 2 = -(b / 2) - 2 := by + rw [← mul_right_inj' (by norm_num : (2 : ℤ) ≠ 0), mul_sub, mul_neg, + eq_sub_of_add_eq (Int.ediv_add_emod _ _), eq_sub_of_add_eq (Int.ediv_add_emod _ _)] + omega + rw [aux (by omega : (-m - 3) + m = (-3 : ℤ)), + aux (by omega : (-m - 2) + (m - 1) = (-3 : ℤ)), + aux (by omega : (-m - 1) + (m - 2) = (-3 : ℤ)), + aux (by omega : -m + (m - 3) = (-3 : ℤ))] + simp only [S_neg_sub_two, Polynomial.eval_neg] + ring_nf + +/-- A formula for $(r_1 r_2)^m$, where $m$ is an integer. -/ +lemma reflection_mul_reflection_zpow (m : ℤ) + (t : R := f y * g x - 2) (ht : t = f y * g x - 2 := by rfl) : + ((reflection hf * reflection hg) ^ m).toLinearMap = + LinearMap.id (R := R) (M := M) + + ((S R ((m - 2) / 2)).eval t * ((S R ((m - 1) / 2)).eval t + (S R ((m - 3) / 2)).eval t)) • + ((g x • f - g).smulRight y - f.smulRight x) + + ((S R ((m - 1) / 2)).eval t * ((S R (m / 2)).eval t + (S R ((m - 2) / 2)).eval t)) • + ((f y • g - f).smulRight x - g.smulRight y) := by + ext z + simpa using reflection_mul_reflection_zpow_apply hf hg m z t ht + +/-- A formula for $(r_1 r_2)^m x$, where $m$ is an integer. This is the special case of +`Module.reflection_mul_reflection_zpow_apply` with $z = x$. -/ +lemma reflection_mul_reflection_zpow_apply_self (m : ℤ) + (t : R := f y * g x - 2) (ht : t = f y * g x - 2 := by rfl) : + ((reflection hf * reflection hg) ^ m) x = + ((S R m).eval t + (S R (m - 1)).eval t) • x + ((S R (m - 1)).eval t * -g x) • y := by + /- Even though this is a special case of `Module.reflection_mul_reflection_zpow_apply`, it is + easier to prove it from scratch. -/ + have S_eval_t_sub_two (k : ℤ) : + (S R (k - 2)).eval t = (f y * g x - 2) * (S R (k - 1)).eval t - (S R k).eval t := by + simp [S_sub_two, ht] + induction m using Int.induction_on with + | hz => simp + | hp m ih => + -- Apply the inductive hypothesis. + rw [add_comm (m : ℤ) 1, zpow_one_add, LinearEquiv.mul_apply, LinearEquiv.mul_apply, ih] + -- Expand out all the reflections and use `hf`, `hg`. + simp only [reflection_apply, map_add, map_sub, map_smul, hf, hg] + -- Equate coefficients of `x` and `y`. + match_scalars + · linear_combination (norm := ring_nf) -S_eval_t_sub_two (m + 1) + · ring_nf + | hn m ih => + -- Apply the inductive hypothesis. + rw [sub_eq_add_neg (-m : ℤ) 1, add_comm (-m : ℤ) (-1), zpow_add, zpow_neg_one, mul_inv_rev, + reflection_inv, reflection_inv, LinearEquiv.mul_apply, LinearEquiv.mul_apply, ih] + -- Expand out all the reflections and use `hf`, `hg`. + simp only [reflection_apply, map_add, map_sub, map_smul, hf, hg] + -- Equate coefficients of `x` and `y`. + match_scalars + · linear_combination (norm := ring_nf) -S_eval_t_sub_two (-m) + · linear_combination (norm := ring_nf) g x * S_eval_t_sub_two (-m) + +/-- A formula for $(r_1 r_2)^m x$, where $m$ is a natural number. This is the special case of +`Module.reflection_mul_reflection_pow_apply` with $z = x$. -/ +lemma reflection_mul_reflection_pow_apply_self (m : ℕ) + (t : R := f y * g x - 2) (ht : t = f y * g x - 2 := by rfl) : + ((reflection hf * reflection hg) ^ m) x = + ((S R m).eval t + (S R (m - 1)).eval t) • x + ((S R (m - 1)).eval t * -g x) • y := + mod_cast reflection_mul_reflection_zpow_apply_self hf hg m t ht + +/-- A formula for $r_2 (r_1 r_2)^m x$, where $m$ is an integer. -/ +lemma reflection_mul_reflection_mul_reflection_zpow_apply_self (m : ℤ) + (t : R := f y * g x - 2) (ht : t = f y * g x - 2 := by rfl) : + (reflection hg * (reflection hf * reflection hg) ^ m) x = + ((S R m).eval t + (S R (m - 1)).eval t) • x + ((S R m).eval t * -g x) • y := by + rw [LinearEquiv.mul_apply, reflection_mul_reflection_zpow_apply_self hf hg m t ht] + -- Expand out all the reflections and use `hf`, `hg`. + simp only [reflection_apply, map_add, map_sub, map_smul, hf, hg] + -- Equate coefficients of `x` and `y`. + module + +/-- A formula for $r_2 (r_1 r_2)^m x$, where $m$ is a natural number. -/ +lemma reflection_mul_reflection_mul_reflection_pow_apply_self (m : ℕ) + (t : R := f y * g x - 2) (ht : t = f y * g x - 2 := by rfl) : + (reflection hg * (reflection hf * reflection hg) ^ m) x = + ((S R m).eval t + (S R (m - 1)).eval t) • x + ((S R m).eval t * -g x) • y := + mod_cast reflection_mul_reflection_mul_reflection_zpow_apply_self hf hg m t ht + +end + +/-! ### Lemmas used to prove uniqueness results for root data -/ + /-- See also `Module.Dual.eq_of_preReflection_mapsTo'` for a variant of this lemma which applies when `Φ` does not span. diff --git a/Mathlib/RingTheory/Polynomial/Chebyshev.lean b/Mathlib/RingTheory/Polynomial/Chebyshev.lean index 3bed452ffb6b3..c795e148a61ce 100644 --- a/Mathlib/RingTheory/Polynomial/Chebyshev.lean +++ b/Mathlib/RingTheory/Polynomial/Chebyshev.lean @@ -372,6 +372,16 @@ theorem S_comp_two_mul_X (n : ℤ) : (S R n).comp (2 * X) = U R n := by | add_two n ih1 ih2 => simp_rw [U_add_two, S_add_two, sub_comp, mul_comp, X_comp, ih1, ih2] | neg_add_one n ih1 ih2 => simp_rw [U_sub_one, S_sub_one, sub_comp, mul_comp, X_comp, ih1, ih2] +theorem S_sq_add_S_sq (n : ℤ) : S R n ^ 2 + S R (n + 1) ^ 2 - X * S R n * S R (n + 1) = 1 := by + induction n using Int.induction_on with + | hz => simp; ring + | hp n ih => + have h₁ := S_add_two R n + linear_combination (norm := ring_nf) (S R (2 + n) - S R n) * h₁ + ih + | hn n ih => + have h₁ := S_sub_one R (-n) + linear_combination (norm := ring_nf) (S R (-1 - n) - S R (1 - n)) * h₁ + ih + theorem S_eq_U_comp_half_mul_X [Invertible (2 : R)] (n : ℤ) : S R n = (U R n).comp (Polynomial.C ⅟2 * X) := by have := congr_arg (·.comp (Polynomial.C ⅟2 * X)) (S_comp_two_mul_X R n) From 8e985737f2c710b9292ba02b93c0c7ce05f980c6 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 6 Dec 2024 06:27:25 +0000 Subject: [PATCH 542/829] chore: bump dependencies (#19758) --- lake-manifest.json | 4 ++-- lakefile.lean | 2 +- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/lake-manifest.json b/lake-manifest.json index 9545da43cc321..31395f1077639 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -45,10 +45,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "8b6048aa0a4a4b6bcf83597802d8dee734e64b7e", + "rev": "43bcb1964528411e47bfa4edd0c87d1face1fce4", "name": "aesop", "manifestFile": "lake-manifest.json", - "inputRev": "v4.15.0-rc1", + "inputRev": "master", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover-community/quote4", diff --git a/lakefile.lean b/lakefile.lean index 27489fbd32314..7f2926ad0e7c5 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -9,7 +9,7 @@ open Lake DSL require "leanprover-community" / "batteries" @ git "main" require "leanprover-community" / "Qq" @ git "v4.14.0" -require "leanprover-community" / "aesop" @ git "v4.15.0-rc1" +require "leanprover-community" / "aesop" @ git "master" require "leanprover-community" / "proofwidgets" @ git "v0.0.48" require "leanprover-community" / "importGraph" @ git "v4.15.0-rc1" require "leanprover-community" / "LeanSearchClient" @ git "main" From ec0985ae8ab60b5bcdca6c8fbc40d3712717e612 Mon Sep 17 00:00:00 2001 From: Jack McKoen Date: Fri, 6 Dec 2024 07:03:15 +0000 Subject: [PATCH 543/829] feat(CategoryTheory): retracts of objects and morphisms (#19233) Defines retracts of objects and morphisms, and basic API. Used in #19135. Co-authored-by: Jack McKoen <104791831+mckoen@users.noreply.github.com> --- Mathlib.lean | 2 + .../MorphismProperty/Retract.lean | 52 +++++++++ Mathlib/CategoryTheory/Retract.lean | 106 ++++++++++++++++++ 3 files changed, 160 insertions(+) create mode 100644 Mathlib/CategoryTheory/MorphismProperty/Retract.lean create mode 100644 Mathlib/CategoryTheory/Retract.lean diff --git a/Mathlib.lean b/Mathlib.lean index 698f1cd681d1f..8ec56685e09e4 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2008,6 +2008,7 @@ import Mathlib.CategoryTheory.MorphismProperty.IsInvertedBy import Mathlib.CategoryTheory.MorphismProperty.Limits import Mathlib.CategoryTheory.MorphismProperty.OverAdjunction import Mathlib.CategoryTheory.MorphismProperty.Representable +import Mathlib.CategoryTheory.MorphismProperty.Retract import Mathlib.CategoryTheory.NatIso import Mathlib.CategoryTheory.NatTrans import Mathlib.CategoryTheory.Noetherian @@ -2046,6 +2047,7 @@ import Mathlib.CategoryTheory.Products.Unitor import Mathlib.CategoryTheory.Quotient import Mathlib.CategoryTheory.Quotient.Linear import Mathlib.CategoryTheory.Quotient.Preadditive +import Mathlib.CategoryTheory.Retract import Mathlib.CategoryTheory.Shift.Basic import Mathlib.CategoryTheory.Shift.CommShift import Mathlib.CategoryTheory.Shift.Induced diff --git a/Mathlib/CategoryTheory/MorphismProperty/Retract.lean b/Mathlib/CategoryTheory/MorphismProperty/Retract.lean new file mode 100644 index 0000000000000..af653b25df92e --- /dev/null +++ b/Mathlib/CategoryTheory/MorphismProperty/Retract.lean @@ -0,0 +1,52 @@ +/- +Copyright (c) 2024 Jack McKoen. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jack McKoen +-/ +import Mathlib.CategoryTheory.Retract +import Mathlib.CategoryTheory.MorphismProperty.Basic + +/-! +# Stability under retracts + +Given `P : MorphismProperty C`, we introduce a typeclass `P.IsStableUnderRetracts` which +is the property that `P` is stable under retracts. + +-/ + +universe v u + +namespace CategoryTheory + +variable {C : Type u} [Category.{v} C] + +namespace MorphismProperty + +/-- A class of morphisms is stable under retracts if a retract of such a morphism still +lies in the class. -/ +class IsStableUnderRetracts (P : MorphismProperty C) : Prop where + of_retract {X Y Z W : C} {f : X ⟶ Y} {g : Z ⟶ W} (h : RetractArrow f g) (hg : P g) : P f + +lemma of_retract {P : MorphismProperty C} [P.IsStableUnderRetracts] + {X Y Z W : C} {f : X ⟶ Y} {g : Z ⟶ W} (h : RetractArrow f g) (hg : P g) : P f := + IsStableUnderRetracts.of_retract h hg + +instance IsStableUnderRetracts.monomorphisms : (monomorphisms C).IsStableUnderRetracts where + of_retract {_ _ _ _ f g} h (hg : Mono g) := ⟨fun α β w ↦ by + rw [← cancel_mono h.i.left, ← cancel_mono g, Category.assoc, Category.assoc, + h.i_w, reassoc_of% w]⟩ + +instance IsStableUnderRetracts.epimorphisms : (epimorphisms C).IsStableUnderRetracts where + of_retract {_ _ _ _ f g} h (hg : Epi g) := ⟨fun α β w ↦ by + rw [← cancel_epi h.r.right, ← cancel_epi g, ← Category.assoc, ← Category.assoc, ← h.r_w, + Category.assoc, Category.assoc, w]⟩ + +instance IsStableUnderRetracts.isomorphisms : (isomorphisms C).IsStableUnderRetracts where + of_retract {X Y Z W f g} h (_ : IsIso _) := by + refine ⟨h.i.right ≫ inv g ≫ h.r.left, ?_, ?_⟩ + · rw [← h.i_w_assoc, IsIso.hom_inv_id_assoc, h.retract_left] + · rw [Category.assoc, Category.assoc, h.r_w, IsIso.inv_hom_id_assoc, h.retract_right] + +end MorphismProperty + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Retract.lean b/Mathlib/CategoryTheory/Retract.lean new file mode 100644 index 0000000000000..952554b768d5a --- /dev/null +++ b/Mathlib/CategoryTheory/Retract.lean @@ -0,0 +1,106 @@ +/- +Copyright (c) 2024 Jack McKoen. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jack McKoen +-/ +import Mathlib.CategoryTheory.Comma.Arrow +import Mathlib.CategoryTheory.EpiMono + +/-! +# Retracts + +Defines retracts of objects and morphisms. + +-/ + +universe v v' u u' + +namespace CategoryTheory + +variable {C : Type u} [Category.{v} C] {D : Type u'} [Category.{v'} D] + +/-- An object `X` is a retract of `Y` if there are morphisms `i : X ⟶ Y` and `r : Y ⟶ X` such +that `i ≫ r = 𝟙 X`. -/ +structure Retract (X Y : C) where + /-- the split monomorphism -/ + i : X ⟶ Y + /-- the split epimorphism -/ + r : Y ⟶ X + retract : i ≫ r = 𝟙 X := by aesop_cat + +namespace Retract + +attribute [reassoc (attr := simp)] retract + +variable {X Y : C} (h : Retract X Y) + +/-- If `X` is a retract of `Y`, then `F.obj X` is a retract of `F.obj Y`. -/ +@[simps] +def map (F : C ⥤ D) : Retract (F.obj X) (F.obj Y) where + i := F.map h.i + r := F.map h.r + retract := by rw [← F.map_comp h.i h.r, h.retract, F.map_id] + +/-- a retract determines a split epimorphism. -/ +@[simps] def splitEpi : SplitEpi h.r where + section_ := h.i + +/-- a retract determines a split monomorphism. -/ +@[simps] def splitMono : SplitMono h.i where + retraction := h.r + +instance : IsSplitEpi h.r := ⟨⟨h.splitEpi⟩⟩ + +instance : IsSplitMono h.i := ⟨⟨h.splitMono⟩⟩ + +end Retract + +/-- +``` + X -------> Z -------> X + | | | + f g f + | | | + v v v + Y -------> W -------> Y + +``` +A morphism `f : X ⟶ Y` is a retract of `g : Z ⟶ W` if there are morphisms `i : f ⟶ g` +and `r : g ⟶ f` in the arrow category such that `i ≫ r = 𝟙 f`. -/ +abbrev RetractArrow {X Y Z W : C} (f : X ⟶ Y) (g : Z ⟶ W) := Retract (Arrow.mk f) (Arrow.mk g) + +namespace RetractArrow + +variable {X Y Z W : C} {f : X ⟶ Y} {g : Z ⟶ W} (h : RetractArrow f g) + +@[reassoc] +lemma i_w : h.i.left ≫ g = f ≫ h.i.right := h.i.w + +@[reassoc] +lemma r_w : h.r.left ≫ f = g ≫ h.r.right := h.r.w + +/-- The top of a retract diagram of morphisms determines a retract of objects. -/ +@[simps!] +def left : Retract X Z := h.map Arrow.leftFunc + +/-- The bottom of a retract diagram of morphisms determines a retract of objects. -/ +@[simps!] +def right : Retract Y W := h.map Arrow.rightFunc + +@[reassoc (attr := simp)] +lemma retract_left : h.i.left ≫ h.r.left = 𝟙 X := h.left.retract + +@[reassoc (attr := simp)] +lemma retract_right : h.i.right ≫ h.r.right = 𝟙 Y := h.right.retract + +instance : IsSplitEpi h.r.left := ⟨⟨h.left.splitEpi⟩⟩ + +instance : IsSplitEpi h.r.right := ⟨⟨h.right.splitEpi⟩⟩ + +instance : IsSplitMono h.i.left := ⟨⟨h.left.splitMono⟩⟩ + +instance : IsSplitMono h.i.right := ⟨⟨h.right.splitMono⟩⟩ + +end RetractArrow + +end CategoryTheory From 6c1905d59194a4bf69ce1f1b2bc04fa3a262f5a8 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Fri, 6 Dec 2024 08:34:42 +0000 Subject: [PATCH 544/829] chore(RingTheory/JacobsonIdeal): split out Polynomial (#19734) Also create a folder RingTheory/Jacobson containing three files Ideal (originally JacobsonIdeal), Ring (original Jacobson), and Polynomial (new). 22 files have transitive imports reduced by 100 or more. --- Mathlib.lean | 5 +- Mathlib/Algebra/CharP/LocalRing.lean | 1 + .../PrimeSpectrum/Jacobson.lean | 2 +- Mathlib/LinearAlgebra/Matrix/Ideal.lean | 2 +- Mathlib/RingTheory/AdicCompletion/Basic.lean | 2 +- Mathlib/RingTheory/Ideal/Over.lean | 2 +- .../Ideal.lean} | 44 +-------------- Mathlib/RingTheory/Jacobson/Polynomial.lean | 54 +++++++++++++++++++ .../{Jacobson.lean => Jacobson/Ring.lean} | 2 +- .../LocalRing/MaximalIdeal/Basic.lean | 2 +- Mathlib/RingTheory/LocalRing/Subring.lean | 3 -- Mathlib/RingTheory/Nakayama.lean | 3 +- Mathlib/RingTheory/Nullstellensatz.lean | 2 +- .../RingTheory/Valuation/LocalSubring.lean | 3 +- 14 files changed, 70 insertions(+), 57 deletions(-) rename Mathlib/RingTheory/{JacobsonIdeal.lean => Jacobson/Ideal.lean} (88%) create mode 100644 Mathlib/RingTheory/Jacobson/Polynomial.lean rename Mathlib/RingTheory/{Jacobson.lean => Jacobson/Ring.lean} (99%) diff --git a/Mathlib.lean b/Mathlib.lean index 8ec56685e09e4..168a5bd80b8dd 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4382,8 +4382,9 @@ import Mathlib.RingTheory.IntegralDomain import Mathlib.RingTheory.IsAdjoinRoot import Mathlib.RingTheory.IsPrimary import Mathlib.RingTheory.IsTensorProduct -import Mathlib.RingTheory.Jacobson -import Mathlib.RingTheory.JacobsonIdeal +import Mathlib.RingTheory.Jacobson.Ideal +import Mathlib.RingTheory.Jacobson.Polynomial +import Mathlib.RingTheory.Jacobson.Ring import Mathlib.RingTheory.Kaehler.Basic import Mathlib.RingTheory.Kaehler.CotangentComplex import Mathlib.RingTheory.Kaehler.Polynomial diff --git a/Mathlib/Algebra/CharP/LocalRing.lean b/Mathlib/Algebra/CharP/LocalRing.lean index e2fb6ba42ddaa..b85b347fd0fd2 100644 --- a/Mathlib/Algebra/CharP/LocalRing.lean +++ b/Mathlib/Algebra/CharP/LocalRing.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Jon Eugster. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jon Eugster -/ +import Mathlib.Algebra.CharP.Defs import Mathlib.Algebra.IsPrimePow import Mathlib.Data.Nat.Factorization.Basic import Mathlib.RingTheory.LocalRing.ResidueField.Defs diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Jacobson.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Jacobson.lean index 57274ae5ae060..f7cf9b2d50040 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Jacobson.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Jacobson.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ import Mathlib.AlgebraicGeometry.PrimeSpectrum.Noetherian -import Mathlib.RingTheory.Jacobson +import Mathlib.RingTheory.Jacobson.Ring import Mathlib.Topology.JacobsonSpace /-! diff --git a/Mathlib/LinearAlgebra/Matrix/Ideal.lean b/Mathlib/LinearAlgebra/Matrix/Ideal.lean index b4de0f855d6dc..ef88aeec52f02 100644 --- a/Mathlib/LinearAlgebra/Matrix/Ideal.lean +++ b/Mathlib/LinearAlgebra/Matrix/Ideal.lean @@ -6,7 +6,7 @@ Authors: Jujian Zhang, Wojciech Nawrocki import Mathlib.Data.Matrix.Basis import Mathlib.RingTheory.Ideal.Lattice import Mathlib.RingTheory.TwoSidedIdeal.Operations -import Mathlib.RingTheory.JacobsonIdeal +import Mathlib.RingTheory.Jacobson.Ideal /-! # Ideals in a matrix ring diff --git a/Mathlib/RingTheory/AdicCompletion/Basic.lean b/Mathlib/RingTheory/AdicCompletion/Basic.lean index a680f5898fdf6..e1dbd4e2cb74f 100644 --- a/Mathlib/RingTheory/AdicCompletion/Basic.lean +++ b/Mathlib/RingTheory/AdicCompletion/Basic.lean @@ -5,7 +5,7 @@ Authors: Kenny Lau, Judith Ludwig, Christian Merten -/ import Mathlib.Algebra.GeomSum import Mathlib.LinearAlgebra.SModEq -import Mathlib.RingTheory.JacobsonIdeal +import Mathlib.RingTheory.Jacobson.Ideal /-! # Completion of a module with respect to an ideal. diff --git a/Mathlib/RingTheory/Ideal/Over.lean b/Mathlib/RingTheory/Ideal/Over.lean index ebb045e9944ca..3442658846e0f 100644 --- a/Mathlib/RingTheory/Ideal/Over.lean +++ b/Mathlib/RingTheory/Ideal/Over.lean @@ -88,7 +88,7 @@ theorem injective_quotient_le_comap_map (P : Ideal R[X]) : R[x] / P → (R / (P ∩ R))[x] / (P / (P ∩ R)) ``` commutes. It is used, for instance, in the proof of `quotient_mk_comp_C_is_integral_of_jacobson`, -in the file `RingTheory.Jacobson`. +in the file `Mathlib.RingTheory.Jacobson.Polynomial`. -/ theorem quotient_mk_maps_eq (P : Ideal R[X]) : ((Quotient.mk (map (mapRingHom (Quotient.mk (P.comap (C : R →+* R[X])))) P)).comp C).comp diff --git a/Mathlib/RingTheory/JacobsonIdeal.lean b/Mathlib/RingTheory/Jacobson/Ideal.lean similarity index 88% rename from Mathlib/RingTheory/JacobsonIdeal.lean rename to Mathlib/RingTheory/Jacobson/Ideal.lean index 804cf08cea5f3..19c3c6e926d8c 100644 --- a/Mathlib/RingTheory/JacobsonIdeal.lean +++ b/Mathlib/RingTheory/Jacobson/Ideal.lean @@ -4,8 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Devon Tuma, Wojciech Nawrocki -/ import Mathlib.RingTheory.Ideal.IsPrimary -import Mathlib.RingTheory.Ideal.Quotient.Basic -import Mathlib.RingTheory.Polynomial.Quotient +import Mathlib.RingTheory.Ideal.Quotient.Operations import Mathlib.RingTheory.TwoSidedIdeal.Operations /-! @@ -52,8 +51,6 @@ namespace Ideal variable {R : Type u} {S : Type v} -open Polynomial - section Jacobson section Ring @@ -338,45 +335,6 @@ end CommRing end Jacobson -section Polynomial - -open Polynomial - -variable [CommRing R] - -theorem jacobson_bot_polynomial_le_sInf_map_maximal : - jacobson (⊥ : Ideal R[X]) ≤ sInf (map (C : R →+* R[X]) '' { J : Ideal R | J.IsMaximal }) := by - refine le_sInf fun J => exists_imp.2 fun j hj => ?_ - haveI : j.IsMaximal := hj.1 - refine Trans.trans (jacobson_mono bot_le) (le_of_eq ?_ : J.jacobson ≤ J) - suffices t : (⊥ : Ideal (Polynomial (R ⧸ j))).jacobson = ⊥ by - rw [← hj.2, jacobson_eq_iff_jacobson_quotient_eq_bot] - replace t := congr_arg (map (polynomialQuotientEquivQuotientPolynomial j).toRingHom) t - rwa [map_jacobson_of_bijective _, map_bot] at t - exact RingEquiv.bijective (polynomialQuotientEquivQuotientPolynomial j) - refine eq_bot_iff.2 fun f hf => ?_ - have r1 : (X : (R ⧸ j)[X]) ≠ 0 := fun hX => by - replace hX := congr_arg (fun f => coeff f 1) hX - simp only [coeff_X_one, coeff_zero] at hX - exact zero_ne_one hX.symm - have r2 := eq_C_of_degree_eq_zero (degree_eq_zero_of_isUnit ((mem_jacobson_bot.1 hf) X)) - simp only [coeff_add, mul_coeff_zero, coeff_X_zero, mul_zero, coeff_one_zero, zero_add] at r2 - erw [add_left_eq_self] at r2 - simpa using (mul_eq_zero.mp r2).resolve_right r1 - -- Porting note: this is golfed to much - -- simpa [(fun hX => by simpa using congr_arg (fun f => coeff f 1) hX : (X : (R ⧸ j)[X]) ≠ 0)] - -- using eq_C_of_degree_eq_zero (degree_eq_zero_of_is_unit ((mem_jacobson_bot.1 hf) X)) - -theorem jacobson_bot_polynomial_of_jacobson_bot (h : jacobson (⊥ : Ideal R) = ⊥) : - jacobson (⊥ : Ideal R[X]) = ⊥ := by - refine eq_bot_iff.2 (le_trans jacobson_bot_polynomial_le_sInf_map_maximal ?_) - refine fun f hf => (Submodule.mem_bot R[X]).2 <| Polynomial.ext fun n => - Trans.trans (?_ : coeff f n = 0) (coeff_zero n).symm - suffices f.coeff n ∈ Ideal.jacobson ⊥ by rwa [h, Submodule.mem_bot] at this - exact mem_sInf.2 fun j hj => (mem_map_C_iff.1 ((mem_sInf.1 hf) ⟨j, ⟨hj.2, rfl⟩⟩)) n - -end Polynomial - section IsLocal variable [CommRing R] diff --git a/Mathlib/RingTheory/Jacobson/Polynomial.lean b/Mathlib/RingTheory/Jacobson/Polynomial.lean new file mode 100644 index 0000000000000..0048db6e6bb7e --- /dev/null +++ b/Mathlib/RingTheory/Jacobson/Polynomial.lean @@ -0,0 +1,54 @@ +/- +Copyright (c) 2020 Devon Tuma. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Devon Tuma +-/ +import Mathlib.RingTheory.Jacobson.Ideal +import Mathlib.RingTheory.Polynomial.Quotient +/-! +# Jacobson radical of polynomial ring + +-/ + +namespace Ideal + +section Polynomial + +open Polynomial + +variable {R : Type*} [CommRing R] + +theorem jacobson_bot_polynomial_le_sInf_map_maximal : + jacobson (⊥ : Ideal R[X]) ≤ sInf (map (C : R →+* R[X]) '' { J : Ideal R | J.IsMaximal }) := by + refine le_sInf fun J => exists_imp.2 fun j hj => ?_ + haveI : j.IsMaximal := hj.1 + refine Trans.trans (jacobson_mono bot_le) (le_of_eq ?_ : J.jacobson ≤ J) + suffices t : (⊥ : Ideal (Polynomial (R ⧸ j))).jacobson = ⊥ by + rw [← hj.2, jacobson_eq_iff_jacobson_quotient_eq_bot] + replace t := congr_arg (map (polynomialQuotientEquivQuotientPolynomial j).toRingHom) t + rwa [map_jacobson_of_bijective _, map_bot] at t + exact RingEquiv.bijective (polynomialQuotientEquivQuotientPolynomial j) + refine eq_bot_iff.2 fun f hf => ?_ + have r1 : (X : (R ⧸ j)[X]) ≠ 0 := fun hX => by + replace hX := congr_arg (fun f => coeff f 1) hX + simp only [coeff_X_one, coeff_zero] at hX + exact zero_ne_one hX.symm + have r2 := eq_C_of_degree_eq_zero (degree_eq_zero_of_isUnit ((mem_jacobson_bot.1 hf) X)) + simp only [coeff_add, mul_coeff_zero, coeff_X_zero, mul_zero, coeff_one_zero, zero_add] at r2 + erw [add_left_eq_self] at r2 + simpa using (mul_eq_zero.mp r2).resolve_right r1 + -- Porting note: this is golfed to much + -- simpa [(fun hX => by simpa using congr_arg (fun f => coeff f 1) hX : (X : (R ⧸ j)[X]) ≠ 0)] + -- using eq_C_of_degree_eq_zero (degree_eq_zero_of_is_unit ((mem_jacobson_bot.1 hf) X)) + +theorem jacobson_bot_polynomial_of_jacobson_bot (h : jacobson (⊥ : Ideal R) = ⊥) : + jacobson (⊥ : Ideal R[X]) = ⊥ := by + refine eq_bot_iff.2 (le_trans jacobson_bot_polynomial_le_sInf_map_maximal ?_) + refine fun f hf => (Submodule.mem_bot R[X]).2 <| Polynomial.ext fun n => + Trans.trans (?_ : coeff f n = 0) (coeff_zero n).symm + suffices f.coeff n ∈ Ideal.jacobson ⊥ by rwa [h, Submodule.mem_bot] at this + exact mem_sInf.2 fun j hj => (mem_map_C_iff.1 ((mem_sInf.1 hf) ⟨j, ⟨hj.2, rfl⟩⟩)) n + +end Polynomial + +end Ideal diff --git a/Mathlib/RingTheory/Jacobson.lean b/Mathlib/RingTheory/Jacobson/Ring.lean similarity index 99% rename from Mathlib/RingTheory/Jacobson.lean rename to Mathlib/RingTheory/Jacobson/Ring.lean index 5fa8e8a0b7c09..a8165786b21e0 100644 --- a/Mathlib/RingTheory/Jacobson.lean +++ b/Mathlib/RingTheory/Jacobson/Ring.lean @@ -5,7 +5,7 @@ Authors: Devon Tuma -/ import Mathlib.RingTheory.Localization.Away.Basic import Mathlib.RingTheory.Ideal.Over -import Mathlib.RingTheory.JacobsonIdeal +import Mathlib.RingTheory.Jacobson.Polynomial import Mathlib.RingTheory.Artinian /-! diff --git a/Mathlib/RingTheory/LocalRing/MaximalIdeal/Basic.lean b/Mathlib/RingTheory/LocalRing/MaximalIdeal/Basic.lean index 2a707385da691..1919f202cbf8f 100644 --- a/Mathlib/RingTheory/LocalRing/MaximalIdeal/Basic.lean +++ b/Mathlib/RingTheory/LocalRing/MaximalIdeal/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Chris Hughes, Mario Carneiro -/ -import Mathlib.RingTheory.JacobsonIdeal +import Mathlib.RingTheory.Jacobson.Ideal import Mathlib.RingTheory.LocalRing.MaximalIdeal.Defs import Mathlib.RingTheory.Localization.Basic import Mathlib.RingTheory.Nilpotent.Lemmas diff --git a/Mathlib/RingTheory/LocalRing/Subring.lean b/Mathlib/RingTheory/LocalRing/Subring.lean index bafa2eb662812..90d074bc764e8 100644 --- a/Mathlib/RingTheory/LocalRing/Subring.lean +++ b/Mathlib/RingTheory/LocalRing/Subring.lean @@ -18,9 +18,6 @@ import Mathlib.RingTheory.Localization.AtPrime open IsLocalRing Set -open scoped Polynomial - - variable {R S : Type*} [CommRing R] [CommRing S] variable {K : Type*} [Field K] diff --git a/Mathlib/RingTheory/Nakayama.lean b/Mathlib/RingTheory/Nakayama.lean index 5d859e8e2189a..05507c2e511a2 100644 --- a/Mathlib/RingTheory/Nakayama.lean +++ b/Mathlib/RingTheory/Nakayama.lean @@ -3,8 +3,9 @@ Copyright (c) 2021 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ +import Mathlib.RingTheory.Finiteness.Basic import Mathlib.RingTheory.Finiteness.Nakayama -import Mathlib.RingTheory.JacobsonIdeal +import Mathlib.RingTheory.Jacobson.Ideal /-! # Nakayama's lemma diff --git a/Mathlib/RingTheory/Nullstellensatz.lean b/Mathlib/RingTheory/Nullstellensatz.lean index 71d381cfeb340..6e41ad765593b 100644 --- a/Mathlib/RingTheory/Nullstellensatz.lean +++ b/Mathlib/RingTheory/Nullstellensatz.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Devon Tuma. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Devon Tuma -/ -import Mathlib.RingTheory.Jacobson +import Mathlib.RingTheory.Jacobson.Ring import Mathlib.FieldTheory.IsAlgClosed.Basic import Mathlib.RingTheory.MvPolynomial import Mathlib.RingTheory.PrimeSpectrum diff --git a/Mathlib/RingTheory/Valuation/LocalSubring.lean b/Mathlib/RingTheory/Valuation/LocalSubring.lean index a3c86e708bac1..ae5fef7208b40 100644 --- a/Mathlib/RingTheory/Valuation/LocalSubring.lean +++ b/Mathlib/RingTheory/Valuation/LocalSubring.lean @@ -3,8 +3,9 @@ Copyright (c) 2024 Andrew Yang, Yaël Dillies, Javier López-Contreras. All righ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang, Yaël Dillies, Javier López-Contreras -/ -import Mathlib.RingTheory.LocalRing.Subring import Mathlib.RingTheory.Ideal.Over +import Mathlib.RingTheory.LocalRing.Subring +import Mathlib.RingTheory.Polynomial.Ideal import Mathlib.RingTheory.Valuation.ValuationSubring /-! From 25a2ed30915f471645506d16d5c0336bd65d3576 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Fri, 6 Dec 2024 09:28:46 +0000 Subject: [PATCH 545/829] chore: address some porting notes for Jacobson ideals (#19752) --- Mathlib/RingTheory/Jacobson/Ideal.lean | 12 +++++------- Mathlib/RingTheory/Jacobson/Polynomial.lean | 13 ++----------- 2 files changed, 7 insertions(+), 18 deletions(-) diff --git a/Mathlib/RingTheory/Jacobson/Ideal.lean b/Mathlib/RingTheory/Jacobson/Ideal.lean index 19c3c6e926d8c..181f68bbf0591 100644 --- a/Mathlib/RingTheory/Jacobson/Ideal.lean +++ b/Mathlib/RingTheory/Jacobson/Ideal.lean @@ -96,8 +96,7 @@ theorem mem_jacobson_iff {x : R} : x ∈ jacobson I ↔ ∀ y, ∃ z, z * y * x let ⟨p, hpi, q, hq, hpq⟩ := Submodule.mem_sup.1 ((eq_top_iff_one _).1 hxy) let ⟨r, hr⟩ := mem_span_singleton'.1 hq ⟨r, by - -- Porting note: supply `mul_add_one` with explicit variables - rw [mul_assoc, ← mul_add_one r (y * x), hr, ← hpq, ← neg_sub, add_sub_cancel_right] + rw [mul_assoc, ← mul_add_one, hr, ← hpq, ← neg_sub, add_sub_cancel_right] exact I.neg_mem hpi⟩) fun hxy : I ⊔ span {y * x + 1} ≠ ⊤ => let ⟨M, hm1, hm2⟩ := exists_le_maximal _ hxy suffices x ∉ M from (this <| mem_sInf.1 hx ⟨le_trans le_sup_left hm2, hm1⟩).elim @@ -108,8 +107,7 @@ theorem mem_jacobson_iff {x : R} : x ∈ jacobson I ↔ ∀ y, ∃ z, z * y * x let ⟨z, hz⟩ := hx (-y) hm.1.1 <| (eq_top_iff_one _).2 <| sub_sub_cancel (z * -y * x + z) 1 ▸ M.sub_mem (by - -- Porting note: supply `mul_add_one` with explicit variables - rw [mul_assoc, ← mul_add_one z, neg_mul, ← sub_eq_iff_eq_add.mpr df.symm, neg_sub, + rw [mul_assoc, ← mul_add_one, neg_mul, ← sub_eq_iff_eq_add.mpr df.symm, neg_sub, sub_add_cancel] exact M.mul_mem_left _ hi) <| him hz⟩ @@ -157,15 +155,15 @@ theorem eq_jacobson_iff_sInf_maximal' : /-- An ideal `I` equals its Jacobson radical if and only if every element outside `I` also lies outside of a maximal ideal containing `I`. -/ theorem eq_jacobson_iff_not_mem : - I.jacobson = I ↔ ∀ (x) (_ : x ∉ I), ∃ M : Ideal R, (I ≤ M ∧ M.IsMaximal) ∧ x ∉ M := by + I.jacobson = I ↔ ∀ x ∉ I, ∃ M : Ideal R, (I ≤ M ∧ M.IsMaximal) ∧ x ∉ M := by constructor · intro h x hx - erw [← h, mem_sInf] at hx + rw [← h, Ideal.jacobson, mem_sInf] at hx push_neg at hx exact hx · refine fun h => le_antisymm (fun x hx => ?_) le_jacobson contrapose hx - erw [mem_sInf] + rw [Ideal.jacobson, mem_sInf] push_neg exact h x hx diff --git a/Mathlib/RingTheory/Jacobson/Polynomial.lean b/Mathlib/RingTheory/Jacobson/Polynomial.lean index 0048db6e6bb7e..c102d0db7133b 100644 --- a/Mathlib/RingTheory/Jacobson/Polynomial.lean +++ b/Mathlib/RingTheory/Jacobson/Polynomial.lean @@ -29,17 +29,8 @@ theorem jacobson_bot_polynomial_le_sInf_map_maximal : rwa [map_jacobson_of_bijective _, map_bot] at t exact RingEquiv.bijective (polynomialQuotientEquivQuotientPolynomial j) refine eq_bot_iff.2 fun f hf => ?_ - have r1 : (X : (R ⧸ j)[X]) ≠ 0 := fun hX => by - replace hX := congr_arg (fun f => coeff f 1) hX - simp only [coeff_X_one, coeff_zero] at hX - exact zero_ne_one hX.symm - have r2 := eq_C_of_degree_eq_zero (degree_eq_zero_of_isUnit ((mem_jacobson_bot.1 hf) X)) - simp only [coeff_add, mul_coeff_zero, coeff_X_zero, mul_zero, coeff_one_zero, zero_add] at r2 - erw [add_left_eq_self] at r2 - simpa using (mul_eq_zero.mp r2).resolve_right r1 - -- Porting note: this is golfed to much - -- simpa [(fun hX => by simpa using congr_arg (fun f => coeff f 1) hX : (X : (R ⧸ j)[X]) ≠ 0)] - -- using eq_C_of_degree_eq_zero (degree_eq_zero_of_is_unit ((mem_jacobson_bot.1 hf) X)) + have r1 : (X : (R ⧸ j)[X]) ≠ 0 := ne_of_apply_ne (coeff · 1) <| by simp + simpa [r1] using eq_C_of_degree_eq_zero (degree_eq_zero_of_isUnit ((mem_jacobson_bot.1 hf) X)) theorem jacobson_bot_polynomial_of_jacobson_bot (h : jacobson (⊥ : Ideal R) = ⊥) : jacobson (⊥ : Ideal R[X]) = ⊥ := by From 80e8a21480592a0555bff6b34d1769eb9db35301 Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Fri, 6 Dec 2024 09:37:03 +0000 Subject: [PATCH 546/829] feat(RingTheory/Flat/Basic): add `Algebra.TensorProduct.include(Left|Right)_injective` (#19745) which states that if certain `algebraMap` is injective and certain module is flat, then `Algebra.TensorProduct.include(Left|Right)` is injective. --- Mathlib/RingTheory/Flat/Basic.lean | 25 +++++++++++++++++++++++++ 1 file changed, 25 insertions(+) diff --git a/Mathlib/RingTheory/Flat/Basic.lean b/Mathlib/RingTheory/Flat/Basic.lean index 9509f2f734fbb..a36eb5785365c 100644 --- a/Mathlib/RingTheory/Flat/Basic.lean +++ b/Mathlib/RingTheory/Flat/Basic.lean @@ -423,3 +423,28 @@ theorem tensorProduct_mapIncl_injective_of_left end Flat end Module + +section Injective + +open scoped TensorProduct + +variable {R S A B : Type*} [CommRing R] [Ring A] [Algebra R A] [Ring B] [Algebra R B] + [CommSemiring S] [Algebra S A] [SMulCommClass R S A] + +namespace Algebra.TensorProduct + +theorem includeLeft_injective [Module.Flat R A] (hb : Function.Injective (algebraMap R B)) : + Function.Injective (includeLeft : A →ₐ[S] A ⊗[R] B) := by + convert Module.Flat.lTensor_preserves_injective_linearMap (M := A) (Algebra.linearMap R B) hb + |>.comp (_root_.TensorProduct.rid R A).symm.injective + ext; simp + +theorem includeRight_injective [Module.Flat R B] (ha : Function.Injective (algebraMap R A)) : + Function.Injective (includeRight : B →ₐ[R] A ⊗[R] B) := by + convert Module.Flat.rTensor_preserves_injective_linearMap (M := B) (Algebra.linearMap R A) ha + |>.comp (_root_.TensorProduct.lid R B).symm.injective + ext; simp + +end Algebra.TensorProduct + +end Injective From ebbff3d1444d7d527d42dca9dfb3e067f8c5e70c Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Fri, 6 Dec 2024 11:40:54 +0000 Subject: [PATCH 547/829] chore(Combinatorics/Hall): remove porting note, others will be fixed by lean#6123 (#19759) --- Mathlib/Combinatorics/Hall/Basic.lean | 7 +++---- 1 file changed, 3 insertions(+), 4 deletions(-) diff --git a/Mathlib/Combinatorics/Hall/Basic.lean b/Mathlib/Combinatorics/Hall/Basic.lean index da15e7bee4a35..b7de187e22a39 100644 --- a/Mathlib/Combinatorics/Hall/Basic.lean +++ b/Mathlib/Combinatorics/Hall/Basic.lean @@ -160,8 +160,7 @@ instance {α : Type u} {β : Type v} [DecidableEq β] (r : α → β → Prop) [∀ a : α, Fintype (Rel.image r {a})] (A : Finset α) : Fintype (Rel.image r A) := by have h : Rel.image r A = (A.biUnion fun a => (Rel.image r {a}).toFinset : Set β) := by ext - -- Porting note: added `Set.mem_toFinset` - simp [Rel.image, (Set.mem_toFinset)] + simp [Rel.image] rw [h] apply FinsetCoe.fintype @@ -185,9 +184,9 @@ theorem Fintype.all_card_le_rel_image_card_iff_exists_injective {α : Type u} { rw [← Set.toFinset_card] apply congr_arg ext b - -- Porting note: added `Set.mem_toFinset` + -- Porting note: added `Set.mem_toFinset` (fixed by lean#6123) simp [Rel.image, (Set.mem_toFinset)] - -- Porting note: added `Set.mem_toFinset` + -- Porting note: added `Set.mem_toFinset` (fixed by lean#6123) have h' : ∀ (f : α → β) (x), r x (f x) ↔ f x ∈ r' x := by simp [Rel.image, (Set.mem_toFinset)] simp only [h, h'] apply Finset.all_card_le_biUnion_card_iff_exists_injective From 997abd13a0b15b5dd161a0822821175db9c6dbbb Mon Sep 17 00:00:00 2001 From: yhtq <1414672068@qq.com> Date: Fri, 6 Dec 2024 13:30:14 +0000 Subject: [PATCH 548/829] fead: Add `lift_unique`, `lift_unique'` to `IsFractionRing` (#19124) Add `lift_unique`, `lift_unique'` to `IsFractionRing` --- Mathlib/RingTheory/Localization/FractionRing.lean | 15 +++++++++++++++ 1 file changed, 15 insertions(+) diff --git a/Mathlib/RingTheory/Localization/FractionRing.lean b/Mathlib/RingTheory/Localization/FractionRing.lean index 137b20d13c55e..309f147722deb 100644 --- a/Mathlib/RingTheory/Localization/FractionRing.lean +++ b/Mathlib/RingTheory/Localization/FractionRing.lean @@ -223,6 +223,21 @@ such that `z = f x * (f y)⁻¹`. -/ noncomputable def lift (hg : Injective g) : K →+* L := IsLocalization.lift fun y : nonZeroDivisors A => isUnit_map_of_injective hg y +theorem lift_unique (hg : Function.Injective g) {f : K →+* L} + (hf1 : ∀ x, f (algebraMap A K x) = g x) : IsFractionRing.lift hg = f := + IsLocalization.lift_unique _ hf1 + +/-- Another version of unique to give two lift maps should be equal -/ +theorem ringHom_ext {f1 f2 : K →+* L} + (hf : ∀ x : A, f1 (algebraMap A K x) = f2 (algebraMap A K x)) : f1 = f2 := by + ext z + obtain ⟨x, y, hy, rfl⟩ := IsFractionRing.div_surjective (A := A) z + rw [map_div₀, map_div₀, hf, hf] + +theorem injective_comp_algebraMap : + Function.Injective fun (f : K →+* L) => f.comp (algebraMap A K) := + fun _ _ h => ringHom_ext (fun x => RingHom.congr_fun h x) + section liftAlgHom variable [Algebra R A] [Algebra R K] [IsScalarTower R A K] [Algebra R L] From e80c9183755c1371c29bd4f5af338f77b8ac2851 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 6 Dec 2024 13:50:00 +0000 Subject: [PATCH 549/829] chore(Filter/Bases): split (#18450) --- Mathlib.lean | 2 + .../MeasureTheory/MeasurableSpace/Basic.lean | 1 + Mathlib/Order/Filter/AtTopBot.lean | 128 ---------- .../Order/Filter/AtTopBot/Archimedean.lean | 1 + .../Filter/AtTopBot/CountablyGenerated.lean | 155 ++++++++++++ Mathlib/Order/Filter/AtTopBot/Field.lean | 2 +- Mathlib/Order/Filter/Bases.lean | 215 ---------------- Mathlib/Order/Filter/Cofinite.lean | 1 + Mathlib/Order/Filter/CountablyGenerated.lean | 233 ++++++++++++++++++ Mathlib/Order/Filter/Defs.lean | 4 + Mathlib/Order/Filter/Pi.lean | 8 - Mathlib/Topology/Bases.lean | 1 + 12 files changed, 399 insertions(+), 352 deletions(-) create mode 100644 Mathlib/Order/Filter/AtTopBot/CountablyGenerated.lean create mode 100644 Mathlib/Order/Filter/CountablyGenerated.lean diff --git a/Mathlib.lean b/Mathlib.lean index 168a5bd80b8dd..9269de2ed24fa 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3995,6 +3995,7 @@ import Mathlib.Order.Extension.Well import Mathlib.Order.Filter.AtTopBot import Mathlib.Order.Filter.AtTopBot.Archimedean import Mathlib.Order.Filter.AtTopBot.BigOperators +import Mathlib.Order.Filter.AtTopBot.CountablyGenerated import Mathlib.Order.Filter.AtTopBot.Field import Mathlib.Order.Filter.AtTopBot.Floor import Mathlib.Order.Filter.AtTopBot.Group @@ -4008,6 +4009,7 @@ import Mathlib.Order.Filter.Cocardinal import Mathlib.Order.Filter.Cofinite import Mathlib.Order.Filter.CountableInter import Mathlib.Order.Filter.CountableSeparatingOn +import Mathlib.Order.Filter.CountablyGenerated import Mathlib.Order.Filter.Curry import Mathlib.Order.Filter.Defs import Mathlib.Order.Filter.ENNReal diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index 9693de1e0a8bf..f0a6159c042b5 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -12,6 +12,7 @@ import Mathlib.GroupTheory.Coset.Defs import Mathlib.MeasureTheory.MeasurableSpace.Instances import Mathlib.Order.Filter.SmallSets import Mathlib.Order.LiminfLimsup +import Mathlib.Order.Filter.AtTopBot.CountablyGenerated import Mathlib.Tactic.FinCases /-! diff --git a/Mathlib/Order/Filter/AtTopBot.lean b/Mathlib/Order/Filter/AtTopBot.lean index ed602722331c0..d8c1fe0ab97b7 100644 --- a/Mathlib/Order/Filter/AtTopBot.lean +++ b/Mathlib/Order/Filter/AtTopBot.lean @@ -172,20 +172,6 @@ theorem Tendsto.eventually_forall_le_atBot [Preorder β] {l : Filter α} ∀ᶠ x in l, ∀ y, y ≤ f x → p y := by rw [← Filter.eventually_forall_le_atBot] at h_evtl; exact (h_evtl.comap f).filter_mono hf.le_comap -instance (priority := 200) atTop.isCountablyGenerated [Preorder α] [Countable α] : - (atTop : Filter <| α).IsCountablyGenerated := - isCountablyGenerated_seq _ - -instance (priority := 200) atBot.isCountablyGenerated [Preorder α] [Countable α] : - (atBot : Filter <| α).IsCountablyGenerated := - isCountablyGenerated_seq _ - -instance _root_.OrderDual.instIsCountablyGeneratedAtTop [Preorder α] - [IsCountablyGenerated (atBot : Filter α)] : IsCountablyGenerated (atTop : Filter αᵒᵈ) := ‹_› - -instance _root_.OrderDual.instIsCountablyGeneratedAtBot [Preorder α] - [IsCountablyGenerated (atTop : Filter α)] : IsCountablyGenerated (atBot : Filter αᵒᵈ) := ‹_› - theorem _root_.IsTop.atTop_eq [Preorder α] {a : α} (ha : IsTop a) : atTop = 𝓟 (Ici a) := (iInf_le _ _).antisymm <| le_iInf fun b ↦ principal_mono.2 <| Ici_subset_Ici.2 <| ha b @@ -291,10 +277,6 @@ theorem map_atTop_eq {f : α → β} : atTop.map f = ⨅ a, 𝓟 (f '' { a' | a theorem frequently_atTop' [NoMaxOrder α] : (∃ᶠ x in atTop, p x) ↔ ∀ a, ∃ b > a, p b := atTop_basis_Ioi.frequently_iff.trans <| by simp -lemma atTop_countable_basis [Countable α] : - HasCountableBasis (atTop : Filter α) (fun _ => True) Ici := - { atTop_basis with countable := to_countable _ } - end IsDirected section IsCodirected @@ -333,10 +315,6 @@ theorem map_atBot_eq {f : α → β} : atBot.map f = ⨅ a, 𝓟 (f '' { a' | a' theorem frequently_atBot' [NoMinOrder α] : (∃ᶠ x in atBot, p x) ↔ ∀ a, ∃ b < a, p b := frequently_atTop' (α := αᵒᵈ) -lemma atBot_countable_basis [Countable α] : - HasCountableBasis (atBot : Filter α) (fun _ => True) Iic := - { atBot_basis with countable := to_countable _ } - end IsCodirected theorem tendsto_atTop [Preorder β] {m : α → β} {f : Filter α} : @@ -918,12 +896,6 @@ theorem prod_atTop_atTop_eq [Preorder α] [Preorder β] : · subsingleton simpa [atTop, prod_iInf_left, prod_iInf_right, iInf_prod] using iInf_comm -instance instIsCountablyGeneratedAtTopProd [Preorder α] [IsCountablyGenerated (atTop : Filter α)] - [Preorder β] [IsCountablyGenerated (atTop : Filter β)] : - IsCountablyGenerated (atTop : Filter (α × β)) := by - rw [← prod_atTop_atTop_eq] - infer_instance - lemma tendsto_finset_prod_atTop : Tendsto (fun (p : Finset ι × Finset ι') ↦ p.1 ×ˢ p.2) atTop atTop := by classical @@ -938,12 +910,6 @@ theorem prod_atBot_atBot_eq [Preorder α] [Preorder β] : (atBot : Filter α) ×ˢ (atBot : Filter β) = (atBot : Filter (α × β)) := @prod_atTop_atTop_eq αᵒᵈ βᵒᵈ _ _ -instance instIsCountablyGeneratedAtBotProd [Preorder α] [IsCountablyGenerated (atBot : Filter α)] - [Preorder β] [IsCountablyGenerated (atBot : Filter β)] : - IsCountablyGenerated (atBot : Filter (α × β)) := by - rw [← prod_atBot_atBot_eq] - infer_instance - theorem prod_map_atTop_eq {α₁ α₂ β₁ β₂ : Type*} [Preorder β₁] [Preorder β₂] (u₁ : β₁ → α₁) (u₂ : β₂ → α₂) : map u₁ atTop ×ˢ map u₂ atTop = map (Prod.map u₁ u₂) atTop := by rw [prod_map_map_eq, prod_atTop_atTop_eq, Prod.map_def] @@ -1296,106 +1262,12 @@ theorem HasAntitoneBasis.subbasis_with_rel {f : Filter α} {s : ℕ → Set α} simp only [forall_mem_image, forall_and, mem_Iio] at hφ exact ⟨φ, forall_swap.2 hφ.1, forall_swap.2 hφ.2⟩ -/-- If `f` is a nontrivial countably generated filter, then there exists a sequence that converges -to `f`. -/ -theorem exists_seq_tendsto (f : Filter α) [IsCountablyGenerated f] [NeBot f] : - ∃ x : ℕ → α, Tendsto x atTop f := by - obtain ⟨B, h⟩ := f.exists_antitone_basis - choose x hx using fun n => Filter.nonempty_of_mem (h.mem n) - exact ⟨x, h.tendsto hx⟩ - -theorem exists_seq_monotone_tendsto_atTop_atTop (α : Type*) [Preorder α] [Nonempty α] - [IsDirected α (· ≤ ·)] [(atTop : Filter α).IsCountablyGenerated] : - ∃ xs : ℕ → α, Monotone xs ∧ Tendsto xs atTop atTop := by - obtain ⟨ys, h⟩ := exists_seq_tendsto (atTop : Filter α) - choose c hleft hright using exists_ge_ge (α := α) - set xs : ℕ → α := fun n => (List.range n).foldl (fun x n ↦ c x (ys n)) (ys 0) - have hsucc (n : ℕ) : xs (n + 1) = c (xs n) (ys n) := by simp [xs, List.range_succ] - refine ⟨xs, ?_, ?_⟩ - · refine monotone_nat_of_le_succ fun n ↦ ?_ - rw [hsucc] - apply hleft - · refine (tendsto_add_atTop_iff_nat 1).1 <| tendsto_atTop_mono (fun n ↦ ?_) h - rw [hsucc] - apply hright - -theorem exists_seq_antitone_tendsto_atTop_atBot (α : Type*) [Preorder α] [Nonempty α] - [IsDirected α (· ≥ ·)] [(atBot : Filter α).IsCountablyGenerated] : - ∃ xs : ℕ → α, Antitone xs ∧ Tendsto xs atTop atBot := - exists_seq_monotone_tendsto_atTop_atTop αᵒᵈ - -/-- An abstract version of continuity of sequentially continuous functions on metric spaces: -if a filter `k` is countably generated then `Tendsto f k l` iff for every sequence `u` -converging to `k`, `f ∘ u` tends to `l`. -/ -theorem tendsto_iff_seq_tendsto {f : α → β} {k : Filter α} {l : Filter β} [k.IsCountablyGenerated] : - Tendsto f k l ↔ ∀ x : ℕ → α, Tendsto x atTop k → Tendsto (f ∘ x) atTop l := by - refine ⟨fun h x hx => h.comp hx, fun H s hs => ?_⟩ - contrapose! H - have : NeBot (k ⊓ 𝓟 (f ⁻¹' sᶜ)) := by simpa [neBot_iff, inf_principal_eq_bot] - rcases (k ⊓ 𝓟 (f ⁻¹' sᶜ)).exists_seq_tendsto with ⟨x, hx⟩ - rw [tendsto_inf, tendsto_principal] at hx - refine ⟨x, hx.1, fun h => ?_⟩ - rcases (hx.2.and (h hs)).exists with ⟨N, hnmem, hmem⟩ - exact hnmem hmem - -theorem tendsto_of_seq_tendsto {f : α → β} {k : Filter α} {l : Filter β} [k.IsCountablyGenerated] : - (∀ x : ℕ → α, Tendsto x atTop k → Tendsto (f ∘ x) atTop l) → Tendsto f k l := - tendsto_iff_seq_tendsto.2 - -theorem eventually_iff_seq_eventually {ι : Type*} {l : Filter ι} {p : ι → Prop} - [l.IsCountablyGenerated] : - (∀ᶠ n in l, p n) ↔ ∀ x : ℕ → ι, Tendsto x atTop l → ∀ᶠ n : ℕ in atTop, p (x n) := by - simpa using tendsto_iff_seq_tendsto (f := id) (l := 𝓟 {x | p x}) - -theorem frequently_iff_seq_frequently {ι : Type*} {l : Filter ι} {p : ι → Prop} - [l.IsCountablyGenerated] : - (∃ᶠ n in l, p n) ↔ ∃ x : ℕ → ι, Tendsto x atTop l ∧ ∃ᶠ n : ℕ in atTop, p (x n) := by - simp only [Filter.Frequently, eventually_iff_seq_eventually (l := l)] - push_neg; rfl - theorem subseq_forall_of_frequently {ι : Type*} {x : ℕ → ι} {p : ι → Prop} {l : Filter ι} (h_tendsto : Tendsto x atTop l) (h : ∃ᶠ n in atTop, p (x n)) : ∃ ns : ℕ → ℕ, Tendsto (fun n => x (ns n)) atTop l ∧ ∀ n, p (x (ns n)) := by choose ns hge hns using frequently_atTop.1 h exact ⟨ns, h_tendsto.comp (tendsto_atTop_mono hge tendsto_id), hns⟩ -theorem exists_seq_forall_of_frequently {ι : Type*} {l : Filter ι} {p : ι → Prop} - [l.IsCountablyGenerated] (h : ∃ᶠ n in l, p n) : - ∃ ns : ℕ → ι, Tendsto ns atTop l ∧ ∀ n, p (ns n) := by - rw [frequently_iff_seq_frequently] at h - obtain ⟨x, hx_tendsto, hx_freq⟩ := h - obtain ⟨n_to_n, h_tendsto, h_freq⟩ := subseq_forall_of_frequently hx_tendsto hx_freq - exact ⟨x ∘ n_to_n, h_tendsto, h_freq⟩ - -lemma frequently_iff_seq_forall {ι : Type*} {l : Filter ι} {p : ι → Prop} - [l.IsCountablyGenerated] : - (∃ᶠ n in l, p n) ↔ ∃ ns : ℕ → ι, Tendsto ns atTop l ∧ ∀ n, p (ns n) := - ⟨exists_seq_forall_of_frequently, fun ⟨_ns, hnsl, hpns⟩ ↦ - hnsl.frequently <| Frequently.of_forall hpns⟩ - -/-- A sequence converges if every subsequence has a convergent subsequence. -/ -theorem tendsto_of_subseq_tendsto {ι : Type*} {x : ι → α} {f : Filter α} {l : Filter ι} - [l.IsCountablyGenerated] - (hxy : ∀ ns : ℕ → ι, Tendsto ns atTop l → - ∃ ms : ℕ → ℕ, Tendsto (fun n => x (ns <| ms n)) atTop f) : - Tendsto x l f := by - contrapose! hxy - obtain ⟨s, hs, hfreq⟩ : ∃ s ∈ f, ∃ᶠ n in l, x n ∉ s := by - rwa [not_tendsto_iff_exists_frequently_nmem] at hxy - obtain ⟨y, hy_tendsto, hy_freq⟩ := exists_seq_forall_of_frequently hfreq - refine ⟨y, hy_tendsto, fun ms hms_tendsto ↦ ?_⟩ - rcases (hms_tendsto.eventually_mem hs).exists with ⟨n, hn⟩ - exact absurd hn <| hy_freq _ - -theorem subseq_tendsto_of_neBot {f : Filter α} [IsCountablyGenerated f] {u : ℕ → α} - (hx : NeBot (f ⊓ map u atTop)) : ∃ θ : ℕ → ℕ, StrictMono θ ∧ Tendsto (u ∘ θ) atTop f := by - rw [← Filter.push_pull', map_neBot_iff] at hx - rcases exists_seq_tendsto (comap u f ⊓ atTop) with ⟨φ, hφ⟩ - rw [tendsto_inf, tendsto_comap_iff] at hφ - obtain ⟨ψ, hψ, hψφ⟩ : ∃ ψ : ℕ → ℕ, StrictMono ψ ∧ StrictMono (φ ∘ ψ) := - strictMono_subseq_of_tendsto_atTop hφ.2 - exact ⟨φ ∘ ψ, hψφ, hφ.1.comp hψ.tendsto_atTop⟩ - end Filter open Filter Finset diff --git a/Mathlib/Order/Filter/AtTopBot/Archimedean.lean b/Mathlib/Order/Filter/AtTopBot/Archimedean.lean index ecc8b33a34a1d..167f93687a47d 100644 --- a/Mathlib/Order/Filter/AtTopBot/Archimedean.lean +++ b/Mathlib/Order/Filter/AtTopBot/Archimedean.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl, Yury Kudryashov -/ import Mathlib.Algebra.Order.Archimedean.Basic import Mathlib.Order.Filter.AtTopBot.Group +import Mathlib.Order.Filter.CountablyGenerated import Mathlib.Tactic.GCongr /-! diff --git a/Mathlib/Order/Filter/AtTopBot/CountablyGenerated.lean b/Mathlib/Order/Filter/AtTopBot/CountablyGenerated.lean new file mode 100644 index 0000000000000..96978e899bf72 --- /dev/null +++ b/Mathlib/Order/Filter/AtTopBot/CountablyGenerated.lean @@ -0,0 +1,155 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Jeremy Avigad, Yury Kudryashov, Patrick Massot +-/ +import Mathlib.Order.Filter.AtTopBot +import Mathlib.Order.Filter.CountablyGenerated + +/-! +# Convergence to infinity and countably generated filters + +In this file we prove that + +- `Filter.atTop` and `Filter.atBot` filters on a countable type are countably generated; +- `Filter.exists_seq_tendsto`: if `f` is a nontrivial countably generated filter, + then there exists a sequence that converges. to `f`; +- `Filter.tendsto_iff_seq_tendsto`: convergence along a countably generated filter + is equivalent to convergence along all sequences that converge to this filter. +-/ + +open Set + +namespace Filter + +variable {α β : Type*} + +instance (priority := 200) atTop.isCountablyGenerated [Preorder α] [Countable α] : + (atTop : Filter <| α).IsCountablyGenerated := + isCountablyGenerated_seq _ + +instance (priority := 200) atBot.isCountablyGenerated [Preorder α] [Countable α] : + (atBot : Filter <| α).IsCountablyGenerated := + isCountablyGenerated_seq _ + +instance instIsCountablyGeneratedAtTopProd [Preorder α] [IsCountablyGenerated (atTop : Filter α)] + [Preorder β] [IsCountablyGenerated (atTop : Filter β)] : + IsCountablyGenerated (atTop : Filter (α × β)) := by + rw [← prod_atTop_atTop_eq] + infer_instance + +instance instIsCountablyGeneratedAtBotProd [Preorder α] [IsCountablyGenerated (atBot : Filter α)] + [Preorder β] [IsCountablyGenerated (atBot : Filter β)] : + IsCountablyGenerated (atBot : Filter (α × β)) := by + rw [← prod_atBot_atBot_eq] + infer_instance + +instance _root_.OrderDual.instIsCountablyGeneratedAtTop [Preorder α] + [IsCountablyGenerated (atBot : Filter α)] : IsCountablyGenerated (atTop : Filter αᵒᵈ) := ‹_› + +instance _root_.OrderDual.instIsCountablyGeneratedAtBot [Preorder α] + [IsCountablyGenerated (atTop : Filter α)] : IsCountablyGenerated (atBot : Filter αᵒᵈ) := ‹_› + +lemma atTop_countable_basis [Preorder α] [IsDirected α (· ≤ ·)] [Nonempty α] [Countable α] : + HasCountableBasis (atTop : Filter α) (fun _ => True) Ici := + { atTop_basis with countable := to_countable _ } + +lemma atBot_countable_basis [Preorder α] [IsDirected α (· ≥ ·)] [Nonempty α] [Countable α] : + HasCountableBasis (atBot : Filter α) (fun _ => True) Iic := + { atBot_basis with countable := to_countable _ } + +/-- If `f` is a nontrivial countably generated filter, then there exists a sequence that converges +to `f`. -/ +theorem exists_seq_tendsto (f : Filter α) [IsCountablyGenerated f] [NeBot f] : + ∃ x : ℕ → α, Tendsto x atTop f := by + obtain ⟨B, h⟩ := f.exists_antitone_basis + choose x hx using fun n => Filter.nonempty_of_mem (h.mem n) + exact ⟨x, h.tendsto hx⟩ + +theorem exists_seq_monotone_tendsto_atTop_atTop (α : Type*) [Preorder α] [Nonempty α] + [IsDirected α (· ≤ ·)] [(atTop : Filter α).IsCountablyGenerated] : + ∃ xs : ℕ → α, Monotone xs ∧ Tendsto xs atTop atTop := by + obtain ⟨ys, h⟩ := exists_seq_tendsto (atTop : Filter α) + choose c hleft hright using exists_ge_ge (α := α) + set xs : ℕ → α := fun n => (List.range n).foldl (fun x n ↦ c x (ys n)) (ys 0) + have hsucc (n : ℕ) : xs (n + 1) = c (xs n) (ys n) := by simp [xs, List.range_succ] + refine ⟨xs, ?_, ?_⟩ + · refine monotone_nat_of_le_succ fun n ↦ ?_ + rw [hsucc] + apply hleft + · refine (tendsto_add_atTop_iff_nat 1).1 <| tendsto_atTop_mono (fun n ↦ ?_) h + rw [hsucc] + apply hright + +theorem exists_seq_antitone_tendsto_atTop_atBot (α : Type*) [Preorder α] [Nonempty α] + [IsDirected α (· ≥ ·)] [(atBot : Filter α).IsCountablyGenerated] : + ∃ xs : ℕ → α, Antitone xs ∧ Tendsto xs atTop atBot := + exists_seq_monotone_tendsto_atTop_atTop αᵒᵈ + +/-- An abstract version of continuity of sequentially continuous functions on metric spaces: +if a filter `k` is countably generated then `Tendsto f k l` iff for every sequence `u` +converging to `k`, `f ∘ u` tends to `l`. -/ +theorem tendsto_iff_seq_tendsto {f : α → β} {k : Filter α} {l : Filter β} [k.IsCountablyGenerated] : + Tendsto f k l ↔ ∀ x : ℕ → α, Tendsto x atTop k → Tendsto (f ∘ x) atTop l := by + refine ⟨fun h x hx => h.comp hx, fun H s hs => ?_⟩ + contrapose! H + have : NeBot (k ⊓ 𝓟 (f ⁻¹' sᶜ)) := by simpa [neBot_iff, inf_principal_eq_bot] + rcases (k ⊓ 𝓟 (f ⁻¹' sᶜ)).exists_seq_tendsto with ⟨x, hx⟩ + rw [tendsto_inf, tendsto_principal] at hx + refine ⟨x, hx.1, fun h => ?_⟩ + rcases (hx.2.and (h hs)).exists with ⟨N, hnmem, hmem⟩ + exact hnmem hmem + +theorem tendsto_of_seq_tendsto {f : α → β} {k : Filter α} {l : Filter β} [k.IsCountablyGenerated] : + (∀ x : ℕ → α, Tendsto x atTop k → Tendsto (f ∘ x) atTop l) → Tendsto f k l := + tendsto_iff_seq_tendsto.2 + +theorem eventually_iff_seq_eventually {ι : Type*} {l : Filter ι} {p : ι → Prop} + [l.IsCountablyGenerated] : + (∀ᶠ n in l, p n) ↔ ∀ x : ℕ → ι, Tendsto x atTop l → ∀ᶠ n : ℕ in atTop, p (x n) := by + simpa using tendsto_iff_seq_tendsto (f := id) (l := 𝓟 {x | p x}) + +theorem frequently_iff_seq_frequently {ι : Type*} {l : Filter ι} {p : ι → Prop} + [l.IsCountablyGenerated] : + (∃ᶠ n in l, p n) ↔ ∃ x : ℕ → ι, Tendsto x atTop l ∧ ∃ᶠ n : ℕ in atTop, p (x n) := by + simp only [Filter.Frequently, eventually_iff_seq_eventually (l := l)] + push_neg; rfl + +theorem exists_seq_forall_of_frequently {ι : Type*} {l : Filter ι} {p : ι → Prop} + [l.IsCountablyGenerated] (h : ∃ᶠ n in l, p n) : + ∃ ns : ℕ → ι, Tendsto ns atTop l ∧ ∀ n, p (ns n) := by + rw [frequently_iff_seq_frequently] at h + obtain ⟨x, hx_tendsto, hx_freq⟩ := h + obtain ⟨n_to_n, h_tendsto, h_freq⟩ := subseq_forall_of_frequently hx_tendsto hx_freq + exact ⟨x ∘ n_to_n, h_tendsto, h_freq⟩ + +lemma frequently_iff_seq_forall {ι : Type*} {l : Filter ι} {p : ι → Prop} + [l.IsCountablyGenerated] : + (∃ᶠ n in l, p n) ↔ ∃ ns : ℕ → ι, Tendsto ns atTop l ∧ ∀ n, p (ns n) := + ⟨exists_seq_forall_of_frequently, fun ⟨_ns, hnsl, hpns⟩ ↦ + hnsl.frequently <| Frequently.of_forall hpns⟩ + +/-- A sequence converges if every subsequence has a convergent subsequence. -/ +theorem tendsto_of_subseq_tendsto {ι : Type*} {x : ι → α} {f : Filter α} {l : Filter ι} + [l.IsCountablyGenerated] + (hxy : ∀ ns : ℕ → ι, Tendsto ns atTop l → + ∃ ms : ℕ → ℕ, Tendsto (fun n => x (ns <| ms n)) atTop f) : + Tendsto x l f := by + contrapose! hxy + obtain ⟨s, hs, hfreq⟩ : ∃ s ∈ f, ∃ᶠ n in l, x n ∉ s := by + rwa [not_tendsto_iff_exists_frequently_nmem] at hxy + obtain ⟨y, hy_tendsto, hy_freq⟩ := exists_seq_forall_of_frequently hfreq + refine ⟨y, hy_tendsto, fun ms hms_tendsto ↦ ?_⟩ + rcases (hms_tendsto.eventually_mem hs).exists with ⟨n, hn⟩ + exact absurd hn <| hy_freq _ + +theorem subseq_tendsto_of_neBot {f : Filter α} [IsCountablyGenerated f] {u : ℕ → α} + (hx : NeBot (f ⊓ map u atTop)) : ∃ θ : ℕ → ℕ, StrictMono θ ∧ Tendsto (u ∘ θ) atTop f := by + rw [← Filter.push_pull', map_neBot_iff] at hx + rcases exists_seq_tendsto (comap u f ⊓ atTop) with ⟨φ, hφ⟩ + rw [tendsto_inf, tendsto_comap_iff] at hφ + obtain ⟨ψ, hψ, hψφ⟩ : ∃ ψ : ℕ → ℕ, StrictMono ψ ∧ StrictMono (φ ∘ ψ) := + strictMono_subseq_of_tendsto_atTop hφ.2 + exact ⟨φ ∘ ψ, hψφ, hφ.1.comp hψ.tendsto_atTop⟩ + +end Filter diff --git a/Mathlib/Order/Filter/AtTopBot/Field.lean b/Mathlib/Order/Filter/AtTopBot/Field.lean index 519009d383724..1112b9f9d9e92 100644 --- a/Mathlib/Order/Filter/AtTopBot/Field.lean +++ b/Mathlib/Order/Filter/AtTopBot/Field.lean @@ -94,7 +94,7 @@ theorem tendsto_const_mul_pow_atTop_iff : exact pos_of_mul_pos_left hck (pow_nonneg hk _) lemma tendsto_zpow_atTop_atTop {n : ℤ} (hn : 0 < n) : Tendsto (fun x : α ↦ x ^ n) atTop atTop := by - lift n to ℕ+ using hn; simp + lift n to ℕ using hn.le; simp [(Int.ofNat_pos.mp hn).ne'] end LinearOrderedSemifield diff --git a/Mathlib/Order/Filter/Bases.lean b/Mathlib/Order/Filter/Bases.lean index ab8dd13872eb6..42dfb3d129084 100644 --- a/Mathlib/Order/Filter/Bases.lean +++ b/Mathlib/Order/Filter/Bases.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Johannes Hölzl, Mario Carneiro, Patrick Massot -/ import Mathlib.Data.Prod.PProd -import Mathlib.Data.Set.Countable import Mathlib.Order.Filter.Finite /-! @@ -77,8 +76,6 @@ machinery, e.g., `simp only [true_and_iff]` or `simp only [forall_const]` can he open Set Filter -section sort - variable {α β γ : Type*} {ι ι' : Sort*} /-- A filter basis `B` on a type `α` is a nonempty collection of sets of `α` @@ -850,215 +847,3 @@ theorem map_sigma_mk_comap {π : α → Type*} {π' : β → Type*} {f : α → apply image_sigmaMk_preimage_sigmaMap hf end Filter - -end sort - -namespace Filter - -variable {α β γ ι : Type*} {ι' : Sort*} - -/-- `IsCountablyGenerated f` means `f = generate s` for some countable `s`. -/ -class IsCountablyGenerated (f : Filter α) : Prop where - /-- There exists a countable set that generates the filter. -/ - out : ∃ s : Set (Set α), s.Countable ∧ f = generate s - -/-- `IsCountableBasis p s` means the image of `s` bounded by `p` is a countable filter basis. -/ -structure IsCountableBasis (p : ι → Prop) (s : ι → Set α) extends IsBasis p s : Prop where - /-- The set of `i` that satisfy the predicate `p` is countable. -/ - countable : (setOf p).Countable - -/-- We say that a filter `l` has a countable basis `s : ι → Set α` bounded by `p : ι → Prop`, -if `t ∈ l` if and only if `t` includes `s i` for some `i` such that `p i`, and the set -defined by `p` is countable. -/ -structure HasCountableBasis (l : Filter α) (p : ι → Prop) (s : ι → Set α) - extends HasBasis l p s : Prop where - /-- The set of `i` that satisfy the predicate `p` is countable. -/ - countable : (setOf p).Countable - -/-- A countable filter basis `B` on a type `α` is a nonempty countable collection of sets of `α` -such that the intersection of two elements of this collection contains some element -of the collection. -/ -structure CountableFilterBasis (α : Type*) extends FilterBasis α where - /-- The set of sets of the filter basis is countable. -/ - countable : sets.Countable - --- For illustration purposes, the countable filter basis defining `(atTop : Filter ℕ)` -instance Nat.inhabitedCountableFilterBasis : Inhabited (CountableFilterBasis ℕ) := - ⟨⟨default, countable_range fun n => Ici n⟩⟩ - -theorem HasCountableBasis.isCountablyGenerated {f : Filter α} {p : ι → Prop} {s : ι → Set α} - (h : f.HasCountableBasis p s) : f.IsCountablyGenerated := - ⟨⟨{ t | ∃ i, p i ∧ s i = t }, h.countable.image s, h.toHasBasis.eq_generate⟩⟩ - -theorem HasBasis.isCountablyGenerated [Countable ι] {f : Filter α} {p : ι → Prop} {s : ι → Set α} - (h : f.HasBasis p s) : f.IsCountablyGenerated := - HasCountableBasis.isCountablyGenerated ⟨h, to_countable _⟩ - -theorem antitone_seq_of_seq (s : ℕ → Set α) : - ∃ t : ℕ → Set α, Antitone t ∧ ⨅ i, 𝓟 (s i) = ⨅ i, 𝓟 (t i) := by - use fun n => ⋂ m ≤ n, s m; constructor - · exact fun i j hij => biInter_mono (Iic_subset_Iic.2 hij) fun n _ => Subset.rfl - apply le_antisymm <;> rw [le_iInf_iff] <;> intro i - · rw [le_principal_iff] - refine (biInter_mem (finite_le_nat _)).2 fun j _ => ?_ - exact mem_iInf_of_mem j (mem_principal_self _) - · refine iInf_le_of_le i (principal_mono.2 <| iInter₂_subset i ?_) - rfl - -theorem countable_biInf_eq_iInf_seq [CompleteLattice α] {B : Set ι} (Bcbl : B.Countable) - (Bne : B.Nonempty) (f : ι → α) : ∃ x : ℕ → ι, ⨅ t ∈ B, f t = ⨅ i, f (x i) := - let ⟨g, hg⟩ := Bcbl.exists_eq_range Bne - ⟨g, hg.symm ▸ iInf_range⟩ - -theorem countable_biInf_eq_iInf_seq' [CompleteLattice α] {B : Set ι} (Bcbl : B.Countable) - (f : ι → α) {i₀ : ι} (h : f i₀ = ⊤) : ∃ x : ℕ → ι, ⨅ t ∈ B, f t = ⨅ i, f (x i) := by - rcases B.eq_empty_or_nonempty with hB | Bnonempty - · rw [hB, iInf_emptyset] - use fun _ => i₀ - simp [h] - · exact countable_biInf_eq_iInf_seq Bcbl Bnonempty f - -theorem countable_biInf_principal_eq_seq_iInf {B : Set (Set α)} (Bcbl : B.Countable) : - ∃ x : ℕ → Set α, ⨅ t ∈ B, 𝓟 t = ⨅ i, 𝓟 (x i) := - countable_biInf_eq_iInf_seq' Bcbl 𝓟 principal_univ - -section IsCountablyGenerated - -protected theorem HasAntitoneBasis.mem_iff [Preorder ι] {l : Filter α} {s : ι → Set α} - (hs : l.HasAntitoneBasis s) {t : Set α} : t ∈ l ↔ ∃ i, s i ⊆ t := - hs.toHasBasis.mem_iff.trans <| by simp only [exists_prop, true_and] - -protected theorem HasAntitoneBasis.mem [Preorder ι] {l : Filter α} {s : ι → Set α} - (hs : l.HasAntitoneBasis s) (i : ι) : s i ∈ l := - hs.toHasBasis.mem_of_mem trivial - -theorem HasAntitoneBasis.hasBasis_ge [Preorder ι] [IsDirected ι (· ≤ ·)] {l : Filter α} - {s : ι → Set α} (hs : l.HasAntitoneBasis s) (i : ι) : l.HasBasis (fun j => i ≤ j) s := - hs.1.to_hasBasis (fun j _ => (exists_ge_ge i j).imp fun _k hk => ⟨hk.1, hs.2 hk.2⟩) fun j _ => - ⟨j, trivial, Subset.rfl⟩ - -/-- If `f` is countably generated and `f.HasBasis p s`, then `f` admits a decreasing basis -enumerated by natural numbers such that all sets have the form `s i`. More precisely, there is a -sequence `i n` such that `p (i n)` for all `n` and `s (i n)` is a decreasing sequence of sets which -forms a basis of `f`-/ -theorem HasBasis.exists_antitone_subbasis {f : Filter α} [h : f.IsCountablyGenerated] - {p : ι' → Prop} {s : ι' → Set α} (hs : f.HasBasis p s) : - ∃ x : ℕ → ι', (∀ i, p (x i)) ∧ f.HasAntitoneBasis fun i => s (x i) := by - obtain ⟨x', hx'⟩ : ∃ x : ℕ → Set α, f = ⨅ i, 𝓟 (x i) := by - rcases h with ⟨s, hsc, rfl⟩ - rw [generate_eq_biInf] - exact countable_biInf_principal_eq_seq_iInf hsc - have : ∀ i, x' i ∈ f := fun i => hx'.symm ▸ (iInf_le (fun i => 𝓟 (x' i)) i) (mem_principal_self _) - let x : ℕ → { i : ι' // p i } := fun n => - Nat.recOn n (hs.index _ <| this 0) fun n xn => - hs.index _ <| inter_mem (this <| n + 1) (hs.mem_of_mem xn.2) - have x_anti : Antitone fun i => s (x i).1 := - antitone_nat_of_succ_le fun i => (hs.set_index_subset _).trans inter_subset_right - have x_subset : ∀ i, s (x i).1 ⊆ x' i := by - rintro (_ | i) - exacts [hs.set_index_subset _, (hs.set_index_subset _).trans inter_subset_left] - refine ⟨fun i => (x i).1, fun i => (x i).2, ?_⟩ - have : (⨅ i, 𝓟 (s (x i).1)).HasAntitoneBasis fun i => s (x i).1 := .iInf_principal x_anti - convert this - exact - le_antisymm (le_iInf fun i => le_principal_iff.2 <| by cases i <;> apply hs.set_index_mem) - (hx'.symm ▸ - le_iInf fun i => le_principal_iff.2 <| this.1.mem_iff.2 ⟨i, trivial, x_subset i⟩) - -/-- A countably generated filter admits a basis formed by an antitone sequence of sets. -/ -theorem exists_antitone_basis (f : Filter α) [f.IsCountablyGenerated] : - ∃ x : ℕ → Set α, f.HasAntitoneBasis x := - let ⟨x, _, hx⟩ := f.basis_sets.exists_antitone_subbasis - ⟨x, hx⟩ - -theorem exists_antitone_seq (f : Filter α) [f.IsCountablyGenerated] : - ∃ x : ℕ → Set α, Antitone x ∧ ∀ {s}, s ∈ f ↔ ∃ i, x i ⊆ s := - let ⟨x, hx⟩ := f.exists_antitone_basis - ⟨x, hx.antitone, by simp [hx.1.mem_iff]⟩ - -instance Inf.isCountablyGenerated (f g : Filter α) [IsCountablyGenerated f] - [IsCountablyGenerated g] : IsCountablyGenerated (f ⊓ g) := by - rcases f.exists_antitone_basis with ⟨s, hs⟩ - rcases g.exists_antitone_basis with ⟨t, ht⟩ - exact HasCountableBasis.isCountablyGenerated ⟨hs.1.inf ht.1, Set.to_countable _⟩ - -instance map.isCountablyGenerated (l : Filter α) [l.IsCountablyGenerated] (f : α → β) : - (map f l).IsCountablyGenerated := - let ⟨_x, hxl⟩ := l.exists_antitone_basis - (hxl.map _).isCountablyGenerated - -instance comap.isCountablyGenerated (l : Filter β) [l.IsCountablyGenerated] (f : α → β) : - (comap f l).IsCountablyGenerated := - let ⟨_x, hxl⟩ := l.exists_antitone_basis - (hxl.comap _).isCountablyGenerated - -instance Sup.isCountablyGenerated (f g : Filter α) [IsCountablyGenerated f] - [IsCountablyGenerated g] : IsCountablyGenerated (f ⊔ g) := by - rcases f.exists_antitone_basis with ⟨s, hs⟩ - rcases g.exists_antitone_basis with ⟨t, ht⟩ - exact HasCountableBasis.isCountablyGenerated ⟨hs.1.sup ht.1, Set.to_countable _⟩ - -instance prod.isCountablyGenerated (la : Filter α) (lb : Filter β) [IsCountablyGenerated la] - [IsCountablyGenerated lb] : IsCountablyGenerated (la ×ˢ lb) := - Filter.Inf.isCountablyGenerated _ _ - -instance coprod.isCountablyGenerated (la : Filter α) (lb : Filter β) [IsCountablyGenerated la] - [IsCountablyGenerated lb] : IsCountablyGenerated (la.coprod lb) := - Filter.Sup.isCountablyGenerated _ _ - -end IsCountablyGenerated - -theorem isCountablyGenerated_seq [Countable ι'] (x : ι' → Set α) : - IsCountablyGenerated (⨅ i, 𝓟 (x i)) := by - use range x, countable_range x - rw [generate_eq_biInf, iInf_range] - -theorem isCountablyGenerated_of_seq {f : Filter α} (h : ∃ x : ℕ → Set α, f = ⨅ i, 𝓟 (x i)) : - f.IsCountablyGenerated := by - rcases h with ⟨x, rfl⟩ - apply isCountablyGenerated_seq - -theorem isCountablyGenerated_biInf_principal {B : Set (Set α)} (h : B.Countable) : - IsCountablyGenerated (⨅ s ∈ B, 𝓟 s) := - isCountablyGenerated_of_seq (countable_biInf_principal_eq_seq_iInf h) - -theorem isCountablyGenerated_iff_exists_antitone_basis {f : Filter α} : - IsCountablyGenerated f ↔ ∃ x : ℕ → Set α, f.HasAntitoneBasis x := by - constructor - · intro h - exact f.exists_antitone_basis - · rintro ⟨x, h⟩ - rw [h.1.eq_iInf] - exact isCountablyGenerated_seq x - -@[instance] -theorem isCountablyGenerated_principal (s : Set α) : IsCountablyGenerated (𝓟 s) := - isCountablyGenerated_of_seq ⟨fun _ => s, iInf_const.symm⟩ - -@[instance] -theorem isCountablyGenerated_pure (a : α) : IsCountablyGenerated (pure a) := by - rw [← principal_singleton] - exact isCountablyGenerated_principal _ - -@[instance] -theorem isCountablyGenerated_bot : IsCountablyGenerated (⊥ : Filter α) := - @principal_empty α ▸ isCountablyGenerated_principal _ - -@[instance] -theorem isCountablyGenerated_top : IsCountablyGenerated (⊤ : Filter α) := - @principal_univ α ▸ isCountablyGenerated_principal _ - --- Porting note: without explicit `Sort u` and `Type v`, Lean 4 uses `ι : Prop` -universe u v - -instance iInf.isCountablyGenerated {ι : Sort u} {α : Type v} [Countable ι] (f : ι → Filter α) - [∀ i, IsCountablyGenerated (f i)] : IsCountablyGenerated (⨅ i, f i) := by - choose s hs using fun i => exists_antitone_basis (f i) - rw [← PLift.down_surjective.iInf_comp] - refine HasCountableBasis.isCountablyGenerated ⟨hasBasis_iInf fun n => (hs _).1, ?_⟩ - refine (countable_range <| Sigma.map ((↑) : Finset (PLift ι) → Set (PLift ι)) fun _ => id).mono ?_ - rintro ⟨I, f⟩ ⟨hI, -⟩ - lift I to Finset (PLift ι) using hI - exact ⟨⟨I, f⟩, rfl⟩ - -end Filter diff --git a/Mathlib/Order/Filter/Cofinite.lean b/Mathlib/Order/Filter/Cofinite.lean index 50fbb43ffe460..abcac7aaded93 100644 --- a/Mathlib/Order/Filter/Cofinite.lean +++ b/Mathlib/Order/Filter/Cofinite.lean @@ -6,6 +6,7 @@ Authors: Johannes Hölzl, Jeremy Avigad, Yury Kudryashov import Mathlib.Data.Finite.Prod import Mathlib.Data.Fintype.Pi import Mathlib.Order.Filter.AtTopBot +import Mathlib.Order.Filter.CountablyGenerated import Mathlib.Order.Filter.Ker import Mathlib.Order.Filter.Pi diff --git a/Mathlib/Order/Filter/CountablyGenerated.lean b/Mathlib/Order/Filter/CountablyGenerated.lean new file mode 100644 index 0000000000000..dd106b65ece2d --- /dev/null +++ b/Mathlib/Order/Filter/CountablyGenerated.lean @@ -0,0 +1,233 @@ +/- +Copyright (c) 2019 Gabriel Ebner. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Gabriel Ebner, Yury Kudryashov, Patrick Massot +-/ +import Mathlib.Data.Set.Countable +import Mathlib.Order.Filter.Bases + +/-! +# Countably generated filters + +In this file we define a typeclass `Filter.IsCountablyGenerated` +saying that a filter is generated by a countable family of sets. + +We also define predicates `Filter.IsCountableBasis` and `Filter.HasCountableBasis` +saying that a specific family of sets is a countable basis. +-/ + +open Set + +namespace Filter + +variable {α β γ ι : Type*} {ι' : Sort*} + +/-- `IsCountablyGenerated f` means `f = generate s` for some countable `s`. -/ +class IsCountablyGenerated (f : Filter α) : Prop where + /-- There exists a countable set that generates the filter. -/ + out : ∃ s : Set (Set α), s.Countable ∧ f = generate s + +/-- `IsCountableBasis p s` means the image of `s` bounded by `p` is a countable filter basis. -/ +structure IsCountableBasis (p : ι → Prop) (s : ι → Set α) extends IsBasis p s : Prop where + /-- The set of `i` that satisfy the predicate `p` is countable. -/ + countable : (setOf p).Countable + +/-- We say that a filter `l` has a countable basis `s : ι → Set α` bounded by `p : ι → Prop`, +if `t ∈ l` if and only if `t` includes `s i` for some `i` such that `p i`, and the set +defined by `p` is countable. -/ +structure HasCountableBasis (l : Filter α) (p : ι → Prop) (s : ι → Set α) + extends HasBasis l p s : Prop where + /-- The set of `i` that satisfy the predicate `p` is countable. -/ + countable : (setOf p).Countable + +/-- A countable filter basis `B` on a type `α` is a nonempty countable collection of sets of `α` +such that the intersection of two elements of this collection contains some element +of the collection. -/ +structure CountableFilterBasis (α : Type*) extends FilterBasis α where + /-- The set of sets of the filter basis is countable. -/ + countable : sets.Countable + +-- For illustration purposes, the countable filter basis defining `(atTop : Filter ℕ)` +instance Nat.inhabitedCountableFilterBasis : Inhabited (CountableFilterBasis ℕ) := + ⟨⟨default, countable_range fun n => Ici n⟩⟩ + +theorem HasCountableBasis.isCountablyGenerated {f : Filter α} {p : ι → Prop} {s : ι → Set α} + (h : f.HasCountableBasis p s) : f.IsCountablyGenerated := + ⟨⟨{ t | ∃ i, p i ∧ s i = t }, h.countable.image s, h.toHasBasis.eq_generate⟩⟩ + +theorem HasBasis.isCountablyGenerated [Countable ι] {f : Filter α} {p : ι → Prop} {s : ι → Set α} + (h : f.HasBasis p s) : f.IsCountablyGenerated := + HasCountableBasis.isCountablyGenerated ⟨h, to_countable _⟩ + +theorem antitone_seq_of_seq (s : ℕ → Set α) : + ∃ t : ℕ → Set α, Antitone t ∧ ⨅ i, 𝓟 (s i) = ⨅ i, 𝓟 (t i) := by + use fun n => ⋂ m ≤ n, s m; constructor + · exact fun i j hij => biInter_mono (Iic_subset_Iic.2 hij) fun n _ => Subset.rfl + apply le_antisymm <;> rw [le_iInf_iff] <;> intro i + · rw [le_principal_iff] + refine (biInter_mem (finite_le_nat _)).2 fun j _ => ?_ + exact mem_iInf_of_mem j (mem_principal_self _) + · refine iInf_le_of_le i (principal_mono.2 <| iInter₂_subset i ?_) + rfl + +theorem countable_biInf_eq_iInf_seq [CompleteLattice α] {B : Set ι} (Bcbl : B.Countable) + (Bne : B.Nonempty) (f : ι → α) : ∃ x : ℕ → ι, ⨅ t ∈ B, f t = ⨅ i, f (x i) := + let ⟨g, hg⟩ := Bcbl.exists_eq_range Bne + ⟨g, hg.symm ▸ iInf_range⟩ + +theorem countable_biInf_eq_iInf_seq' [CompleteLattice α] {B : Set ι} (Bcbl : B.Countable) + (f : ι → α) {i₀ : ι} (h : f i₀ = ⊤) : ∃ x : ℕ → ι, ⨅ t ∈ B, f t = ⨅ i, f (x i) := by + rcases B.eq_empty_or_nonempty with hB | Bnonempty + · rw [hB, iInf_emptyset] + use fun _ => i₀ + simp [h] + · exact countable_biInf_eq_iInf_seq Bcbl Bnonempty f + +theorem countable_biInf_principal_eq_seq_iInf {B : Set (Set α)} (Bcbl : B.Countable) : + ∃ x : ℕ → Set α, ⨅ t ∈ B, 𝓟 t = ⨅ i, 𝓟 (x i) := + countable_biInf_eq_iInf_seq' Bcbl 𝓟 principal_univ + +section IsCountablyGenerated + +protected theorem HasAntitoneBasis.mem_iff [Preorder ι] {l : Filter α} {s : ι → Set α} + (hs : l.HasAntitoneBasis s) {t : Set α} : t ∈ l ↔ ∃ i, s i ⊆ t := + hs.toHasBasis.mem_iff.trans <| by simp only [exists_prop, true_and] + +protected theorem HasAntitoneBasis.mem [Preorder ι] {l : Filter α} {s : ι → Set α} + (hs : l.HasAntitoneBasis s) (i : ι) : s i ∈ l := + hs.toHasBasis.mem_of_mem trivial + +theorem HasAntitoneBasis.hasBasis_ge [Preorder ι] [IsDirected ι (· ≤ ·)] {l : Filter α} + {s : ι → Set α} (hs : l.HasAntitoneBasis s) (i : ι) : l.HasBasis (fun j => i ≤ j) s := + hs.1.to_hasBasis (fun j _ => (exists_ge_ge i j).imp fun _k hk => ⟨hk.1, hs.2 hk.2⟩) fun j _ => + ⟨j, trivial, Subset.rfl⟩ + +/-- If `f` is countably generated and `f.HasBasis p s`, then `f` admits a decreasing basis +enumerated by natural numbers such that all sets have the form `s i`. More precisely, there is a +sequence `i n` such that `p (i n)` for all `n` and `s (i n)` is a decreasing sequence of sets which +forms a basis of `f`-/ +theorem HasBasis.exists_antitone_subbasis {f : Filter α} [h : f.IsCountablyGenerated] + {p : ι' → Prop} {s : ι' → Set α} (hs : f.HasBasis p s) : + ∃ x : ℕ → ι', (∀ i, p (x i)) ∧ f.HasAntitoneBasis fun i => s (x i) := by + obtain ⟨x', hx'⟩ : ∃ x : ℕ → Set α, f = ⨅ i, 𝓟 (x i) := by + rcases h with ⟨s, hsc, rfl⟩ + rw [generate_eq_biInf] + exact countable_biInf_principal_eq_seq_iInf hsc + have : ∀ i, x' i ∈ f := fun i => hx'.symm ▸ (iInf_le (fun i => 𝓟 (x' i)) i) (mem_principal_self _) + let x : ℕ → { i : ι' // p i } := fun n => + Nat.recOn n (hs.index _ <| this 0) fun n xn => + hs.index _ <| inter_mem (this <| n + 1) (hs.mem_of_mem xn.2) + have x_anti : Antitone fun i => s (x i).1 := + antitone_nat_of_succ_le fun i => (hs.set_index_subset _).trans inter_subset_right + have x_subset : ∀ i, s (x i).1 ⊆ x' i := by + rintro (_ | i) + exacts [hs.set_index_subset _, (hs.set_index_subset _).trans inter_subset_left] + refine ⟨fun i => (x i).1, fun i => (x i).2, ?_⟩ + have : (⨅ i, 𝓟 (s (x i).1)).HasAntitoneBasis fun i => s (x i).1 := .iInf_principal x_anti + convert this + exact + le_antisymm (le_iInf fun i => le_principal_iff.2 <| by cases i <;> apply hs.set_index_mem) + (hx'.symm ▸ + le_iInf fun i => le_principal_iff.2 <| this.1.mem_iff.2 ⟨i, trivial, x_subset i⟩) + +/-- A countably generated filter admits a basis formed by an antitone sequence of sets. -/ +theorem exists_antitone_basis (f : Filter α) [f.IsCountablyGenerated] : + ∃ x : ℕ → Set α, f.HasAntitoneBasis x := + let ⟨x, _, hx⟩ := f.basis_sets.exists_antitone_subbasis + ⟨x, hx⟩ + +theorem exists_antitone_seq (f : Filter α) [f.IsCountablyGenerated] : + ∃ x : ℕ → Set α, Antitone x ∧ ∀ {s}, s ∈ f ↔ ∃ i, x i ⊆ s := + let ⟨x, hx⟩ := f.exists_antitone_basis + ⟨x, hx.antitone, by simp [hx.1.mem_iff]⟩ + +instance Inf.isCountablyGenerated (f g : Filter α) [IsCountablyGenerated f] + [IsCountablyGenerated g] : IsCountablyGenerated (f ⊓ g) := by + rcases f.exists_antitone_basis with ⟨s, hs⟩ + rcases g.exists_antitone_basis with ⟨t, ht⟩ + exact HasCountableBasis.isCountablyGenerated ⟨hs.1.inf ht.1, Set.to_countable _⟩ + +instance map.isCountablyGenerated (l : Filter α) [l.IsCountablyGenerated] (f : α → β) : + (map f l).IsCountablyGenerated := + let ⟨_x, hxl⟩ := l.exists_antitone_basis + (hxl.map _).isCountablyGenerated + +instance comap.isCountablyGenerated (l : Filter β) [l.IsCountablyGenerated] (f : α → β) : + (comap f l).IsCountablyGenerated := + let ⟨_x, hxl⟩ := l.exists_antitone_basis + (hxl.comap _).isCountablyGenerated + +instance Sup.isCountablyGenerated (f g : Filter α) [IsCountablyGenerated f] + [IsCountablyGenerated g] : IsCountablyGenerated (f ⊔ g) := by + rcases f.exists_antitone_basis with ⟨s, hs⟩ + rcases g.exists_antitone_basis with ⟨t, ht⟩ + exact HasCountableBasis.isCountablyGenerated ⟨hs.1.sup ht.1, Set.to_countable _⟩ + +instance prod.isCountablyGenerated (la : Filter α) (lb : Filter β) [IsCountablyGenerated la] + [IsCountablyGenerated lb] : IsCountablyGenerated (la ×ˢ lb) := + Filter.Inf.isCountablyGenerated _ _ + +instance coprod.isCountablyGenerated (la : Filter α) (lb : Filter β) [IsCountablyGenerated la] + [IsCountablyGenerated lb] : IsCountablyGenerated (la.coprod lb) := + Filter.Sup.isCountablyGenerated _ _ + +end IsCountablyGenerated + +theorem isCountablyGenerated_seq [Countable ι'] (x : ι' → Set α) : + IsCountablyGenerated (⨅ i, 𝓟 (x i)) := by + use range x, countable_range x + rw [generate_eq_biInf, iInf_range] + +theorem isCountablyGenerated_of_seq {f : Filter α} (h : ∃ x : ℕ → Set α, f = ⨅ i, 𝓟 (x i)) : + f.IsCountablyGenerated := by + rcases h with ⟨x, rfl⟩ + apply isCountablyGenerated_seq + +theorem isCountablyGenerated_biInf_principal {B : Set (Set α)} (h : B.Countable) : + IsCountablyGenerated (⨅ s ∈ B, 𝓟 s) := + isCountablyGenerated_of_seq (countable_biInf_principal_eq_seq_iInf h) + +theorem isCountablyGenerated_iff_exists_antitone_basis {f : Filter α} : + IsCountablyGenerated f ↔ ∃ x : ℕ → Set α, f.HasAntitoneBasis x := by + constructor + · intro h + exact f.exists_antitone_basis + · rintro ⟨x, h⟩ + rw [h.1.eq_iInf] + exact isCountablyGenerated_seq x + +@[instance] +theorem isCountablyGenerated_principal (s : Set α) : IsCountablyGenerated (𝓟 s) := + isCountablyGenerated_of_seq ⟨fun _ => s, iInf_const.symm⟩ + +@[instance] +theorem isCountablyGenerated_pure (a : α) : IsCountablyGenerated (pure a) := by + rw [← principal_singleton] + exact isCountablyGenerated_principal _ + +@[instance] +theorem isCountablyGenerated_bot : IsCountablyGenerated (⊥ : Filter α) := + @principal_empty α ▸ isCountablyGenerated_principal _ + +@[instance] +theorem isCountablyGenerated_top : IsCountablyGenerated (⊤ : Filter α) := + @principal_univ α ▸ isCountablyGenerated_principal _ + +-- Porting note: without explicit `Sort u` and `Type v`, Lean 4 uses `ι : Prop` +universe u v + +instance iInf.isCountablyGenerated {ι : Sort u} {α : Type v} [Countable ι] (f : ι → Filter α) + [∀ i, IsCountablyGenerated (f i)] : IsCountablyGenerated (⨅ i, f i) := by + choose s hs using fun i => exists_antitone_basis (f i) + rw [← PLift.down_surjective.iInf_comp] + refine HasCountableBasis.isCountablyGenerated ⟨hasBasis_iInf fun n => (hs _).1, ?_⟩ + refine (countable_range <| Sigma.map ((↑) : Finset (PLift ι) → Set (PLift ι)) fun _ => id).mono ?_ + rintro ⟨I, f⟩ ⟨hI, -⟩ + lift I to Finset (PLift ι) using hI + exact ⟨⟨I, f⟩, rfl⟩ + +instance pi.isCountablyGenerated {ι : Type*} {α : ι → Type*} [Countable ι] + (f : ∀ i, Filter (α i)) [∀ i, IsCountablyGenerated (f i)] : IsCountablyGenerated (pi f) := + iInf.isCountablyGenerated _ + +end Filter diff --git a/Mathlib/Order/Filter/Defs.lean b/Mathlib/Order/Filter/Defs.lean index 432915dcbac1d..769d8eb31e206 100644 --- a/Mathlib/Order/Filter/Defs.lean +++ b/Mathlib/Order/Filter/Defs.lean @@ -319,6 +319,10 @@ protected def prod (f : Filter α) (g : Filter β) : Filter (α × β) := f ×ˢ theorem prod_eq_inf (f : Filter α) (g : Filter β) : f ×ˢ g = f.comap Prod.fst ⊓ g.comap Prod.snd := rfl +/-- The product of an indexed family of filters. -/ +def pi {ι : Type*} {α : ι → Type*} (f : ∀ i, Filter (α i)) : Filter (∀ i, α i) := + ⨅ i, comap (Function.eval i) (f i) + /-- The monadic bind operation on filter is defined the usual way in terms of `map` and `join`. Unfortunately, this `bind` does not result in the expected applicative. See `Filter.seq` for the diff --git a/Mathlib/Order/Filter/Pi.lean b/Mathlib/Order/Filter/Pi.lean index 10716f6ff6edc..e583584b9123e 100644 --- a/Mathlib/Order/Filter/Pi.lean +++ b/Mathlib/Order/Filter/Pi.lean @@ -30,14 +30,6 @@ variable {ι : Type*} {α : ι → Type*} {f f₁ f₂ : (i : ι) → Filter (α section Pi -/-- The product of an indexed family of filters. -/ -def pi (f : ∀ i, Filter (α i)) : Filter (∀ i, α i) := - ⨅ i, comap (eval i) (f i) - -instance pi.isCountablyGenerated [Countable ι] [∀ i, IsCountablyGenerated (f i)] : - IsCountablyGenerated (pi f) := - iInf.isCountablyGenerated _ - theorem tendsto_eval_pi (f : ∀ i, Filter (α i)) (i : ι) : Tendsto (eval i) (pi f) (f i) := tendsto_iInf' i tendsto_comap diff --git a/Mathlib/Topology/Bases.lean b/Mathlib/Topology/Bases.lean index 4db6b21371950..969cd3b6dd3fa 100644 --- a/Mathlib/Topology/Bases.lean +++ b/Mathlib/Topology/Bases.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ import Mathlib.Data.Set.Constructions +import Mathlib.Order.Filter.AtTopBot.CountablyGenerated import Mathlib.Topology.Constructions import Mathlib.Topology.ContinuousOn From 6581d46be3ac8d9cbf55ad0cbbfcaacf5e16081f Mon Sep 17 00:00:00 2001 From: Daniel Weber Date: Fri, 6 Dec 2024 14:50:53 +0000 Subject: [PATCH 550/829] chore: eliminate `multiplicity` namespace (#19623) --- Archive/Imo/Imo2019Q4.lean | 2 +- .../Wiedijk100Theorems/PerfectNumbers.lean | 2 +- Mathlib/Algebra/Polynomial/Div.lean | 24 +- Mathlib/Algebra/Polynomial/RingDivision.lean | 2 +- Mathlib/Algebra/Squarefree/Basic.lean | 24 +- Mathlib/Data/Nat/Choose/Factorization.lean | 4 +- Mathlib/Data/Nat/Factorization/Basic.lean | 2 +- Mathlib/Data/Nat/Multiplicity.lean | 16 +- Mathlib/Data/Nat/Squarefree.lean | 2 +- Mathlib/Data/Real/Irrational.lean | 4 +- Mathlib/FieldTheory/Separable.lean | 4 +- Mathlib/FieldTheory/SeparableDegree.lean | 2 +- Mathlib/NumberTheory/FLT/Three.lean | 16 +- Mathlib/NumberTheory/Multiplicity.lean | 78 ++-- Mathlib/NumberTheory/Padics/PadicNorm.lean | 4 +- Mathlib/NumberTheory/Padics/PadicNumbers.lean | 2 +- .../NumberTheory/Padics/PadicVal/Basic.lean | 22 +- .../NumberTheory/Padics/PadicVal/Defs.lean | 12 +- Mathlib/RingTheory/ChainOfDivisors.lean | 8 +- Mathlib/RingTheory/DedekindDomain/Ideal.lean | 12 +- Mathlib/RingTheory/Multiplicity.lean | 334 +++++++++++++----- Mathlib/RingTheory/PowerSeries/Order.lean | 2 - Mathlib/RingTheory/RootsOfUnity/Minpoly.lean | 2 +- .../RootsOfUnity/PrimitiveRoots.lean | 2 +- .../Multiplicity.lean | 14 +- 25 files changed, 373 insertions(+), 223 deletions(-) diff --git a/Archive/Imo/Imo2019Q4.lean b/Archive/Imo/Imo2019Q4.lean index fe10780d6586f..1bcfbe0c77878 100644 --- a/Archive/Imo/Imo2019Q4.lean +++ b/Archive/Imo/Imo2019Q4.lean @@ -30,7 +30,7 @@ open scoped Nat open Nat hiding zero_le Prime -open Finset multiplicity +open Finset namespace Imo2019Q4 diff --git a/Archive/Wiedijk100Theorems/PerfectNumbers.lean b/Archive/Wiedijk100Theorems/PerfectNumbers.lean index a97ff90b5c744..d9d5bb09b3335 100644 --- a/Archive/Wiedijk100Theorems/PerfectNumbers.lean +++ b/Archive/Wiedijk100Theorems/PerfectNumbers.lean @@ -55,7 +55,7 @@ theorem even_two_pow_mul_mersenne_of_prime (k : ℕ) (pr : (mersenne (k + 1)).Pr Even (2 ^ k * mersenne (k + 1)) := by simp [ne_zero_of_prime_mersenne k pr, parity_simps] theorem eq_two_pow_mul_odd {n : ℕ} (hpos : 0 < n) : ∃ k m : ℕ, n = 2 ^ k * m ∧ ¬Even m := by - have h := Nat.multiplicity_finite_iff.2 ⟨Nat.prime_two.ne_one, hpos⟩ + have h := Nat.finiteMultiplicity_iff.2 ⟨Nat.prime_two.ne_one, hpos⟩ cases' pow_multiplicity_dvd 2 n with m hm use multiplicity 2 n, m refine ⟨hm, ?_⟩ diff --git a/Mathlib/Algebra/Polynomial/Div.lean b/Mathlib/Algebra/Polynomial/Div.lean index 42daf78354251..41009d47c194c 100644 --- a/Mathlib/Algebra/Polynomial/Div.lean +++ b/Mathlib/Algebra/Polynomial/Div.lean @@ -53,8 +53,8 @@ theorem X_pow_dvd_iff {f : R[X]} {n : ℕ} : X ^ n ∣ f ↔ ∀ d < n, f.coeff variable {p q : R[X]} -theorem multiplicity_finite_of_degree_pos_of_monic (hp : (0 : WithBot ℕ) < degree p) (hmp : Monic p) - (hq : q ≠ 0) : multiplicity.Finite p q := +theorem finiteMultiplicity_of_degree_pos_of_monic (hp : (0 : WithBot ℕ) < degree p) (hmp : Monic p) + (hq : q ≠ 0) : FiniteMultiplicity p q := have zn0 : (0 : R) ≠ 1 := haveI := Nontrivial.of_polynomial_ne hq zero_ne_one @@ -76,6 +76,9 @@ theorem multiplicity_finite_of_degree_pos_of_monic (hp : (0 : WithBot ℕ) < deg (add_pos_of_pos_of_nonneg (by rwa [one_mul]) (Nat.zero_le _))) this⟩ +@[deprecated (since := "2024-11-30")] +alias multiplicity_finite_of_degree_pos_of_monic := finiteMultiplicity_of_degree_pos_of_monic + end Semiring section Ring @@ -470,12 +473,15 @@ See `polynomial.modByMonic` for the algorithm that computes `%ₘ`. def decidableDvdMonic [DecidableEq R] (p : R[X]) (hq : Monic q) : Decidable (q ∣ p) := decidable_of_iff (p %ₘ q = 0) (modByMonic_eq_zero_iff_dvd hq) -theorem multiplicity_X_sub_C_finite (a : R) (h0 : p ≠ 0) : multiplicity.Finite (X - C a) p := by +theorem finiteMultiplicity_X_sub_C (a : R) (h0 : p ≠ 0) : FiniteMultiplicity (X - C a) p := by haveI := Nontrivial.of_polynomial_ne h0 - refine multiplicity_finite_of_degree_pos_of_monic ?_ (monic_X_sub_C _) h0 + refine finiteMultiplicity_of_degree_pos_of_monic ?_ (monic_X_sub_C _) h0 rw [degree_X_sub_C] decide +@[deprecated (since := "2024-11-30")] +alias multiplicity_X_sub_C_finite := finiteMultiplicity_X_sub_C + /- Porting note: stripping out classical for decidability instance parameter might make for better ergonomics -/ /-- The largest power of `X - C a` which divides `p`. @@ -488,7 +494,7 @@ def rootMultiplicity (a : R) (p : R[X]) : ℕ := let _ : DecidablePred fun n : ℕ => ¬(X - C a) ^ (n + 1) ∣ p := fun n => have := decidableDvdMonic p ((monic_X_sub_C a).pow (n + 1)) inferInstanceAs (Decidable ¬_) - Nat.find (multiplicity_X_sub_C_finite a h0) + Nat.find (finiteMultiplicity_X_sub_C a h0) /- Porting note: added the following due to diamond with decidableProp and decidableDvdMonic see also [Zulip] @@ -497,7 +503,7 @@ theorem rootMultiplicity_eq_nat_find_of_nonzero [DecidableEq R] {p : R[X]} (p0 : letI : DecidablePred fun n : ℕ => ¬(X - C a) ^ (n + 1) ∣ p := fun n => have := decidableDvdMonic p ((monic_X_sub_C a).pow (n + 1)) inferInstanceAs (Decidable ¬_) - rootMultiplicity a p = Nat.find (multiplicity_X_sub_C_finite a p0) := by + rootMultiplicity a p = Nat.find (finiteMultiplicity_X_sub_C a p0) := by dsimp [rootMultiplicity] cases Subsingleton.elim ‹DecidableEq R› (Classical.decEq R) rw [dif_neg p0] @@ -510,7 +516,7 @@ theorem rootMultiplicity_eq_multiplicity [DecidableEq R] split · rfl rename_i h - simp only [multiplicity_X_sub_C_finite a h, ↓reduceDIte] + simp only [finiteMultiplicity_X_sub_C a h, ↓reduceDIte] rw [← ENat.some_eq_coe, WithTop.untop'_coe] congr @@ -548,7 +554,7 @@ theorem exists_eq_pow_rootMultiplicity_mul_and_not_dvd (p : R[X]) (hp : p ≠ 0) ∃ q : R[X], p = (X - C a) ^ p.rootMultiplicity a * q ∧ ¬ (X - C a) ∣ q := by classical rw [rootMultiplicity_eq_multiplicity, if_neg hp] - apply (multiplicity_X_sub_C_finite a hp).exists_eq_pow_mul_and_not_dvd + apply (finiteMultiplicity_X_sub_C a hp).exists_eq_pow_mul_and_not_dvd end multiplicity @@ -633,7 +639,7 @@ theorem eval_divByMonic_pow_rootMultiplicity_ne_zero {p : R[X]} (a : R) (hp : p have := pow_mul_divByMonic_rootMultiplicity_eq p a rw [hq, ← mul_assoc, ← pow_succ, rootMultiplicity_eq_multiplicity, if_neg hp] at this exact - (multiplicity_finite_of_degree_pos_of_monic + (finiteMultiplicity_of_degree_pos_of_monic (show (0 : WithBot ℕ) < degree (X - C a) by rw [degree_X_sub_C]; decide) (monic_X_sub_C _) hp).not_pow_dvd_of_multiplicity_lt (Nat.lt_succ_self _) (dvd_of_mul_right_eq _ this) diff --git a/Mathlib/Algebra/Polynomial/RingDivision.lean b/Mathlib/Algebra/Polynomial/RingDivision.lean index b8bd2125c5b79..ad3c8e6a0e543 100644 --- a/Mathlib/Algebra/Polynomial/RingDivision.lean +++ b/Mathlib/Algebra/Polynomial/RingDivision.lean @@ -230,7 +230,7 @@ theorem rootMultiplicity_mul {p q : R[X]} {x : R} (hpq : p * q ≠ 0) : have hq : q ≠ 0 := right_ne_zero_of_mul hpq rw [rootMultiplicity_eq_multiplicity (p * q), if_neg hpq, rootMultiplicity_eq_multiplicity p, if_neg hp, rootMultiplicity_eq_multiplicity q, if_neg hq, - multiplicity_mul (prime_X_sub_C x) (multiplicity_X_sub_C_finite _ hpq)] + multiplicity_mul (prime_X_sub_C x) (finiteMultiplicity_X_sub_C _ hpq)] open Multiset in set_option linter.unusedVariables false in diff --git a/Mathlib/Algebra/Squarefree/Basic.lean b/Mathlib/Algebra/Squarefree/Basic.lean index 9d68d11fc173a..bdc6c55ec4a16 100644 --- a/Mathlib/Algebra/Squarefree/Basic.lean +++ b/Mathlib/Algebra/Squarefree/Basic.lean @@ -103,13 +103,7 @@ theorem Squarefree.gcd_left {a : α} (b : α) (ha : Squarefree a) : Squarefree ( end SquarefreeGcdOfSquarefree -namespace multiplicity - -section CommMonoid - -variable [CommMonoid R] - -theorem squarefree_iff_emultiplicity_le_one (r : R) : +theorem squarefree_iff_emultiplicity_le_one [CommMonoid R] (r : R) : Squarefree r ↔ ∀ x : R, emultiplicity x r ≤ 1 ∨ IsUnit x := by refine forall_congr' fun a => ?_ rw [← sq, pow_dvd_iff_le_emultiplicity, or_iff_not_imp_left, not_le, imp_congr _ Iff.rfl] @@ -117,18 +111,8 @@ theorem squarefree_iff_emultiplicity_le_one (r : R) : rw [← one_add_one_eq_two] exact Order.add_one_le_iff_of_not_isMax (by simp) -end CommMonoid - -section CancelCommMonoidWithZero - -variable [CancelCommMonoidWithZero R] [WfDvdMonoid R] - -theorem finite_prime_left {a b : R} (ha : Prime a) (hb : b ≠ 0) : multiplicity.Finite a b := - finite_of_not_isUnit ha.not_unit hb - -end CancelCommMonoidWithZero - -end multiplicity +@[deprecated (since := "2024-11-30")] +alias multiplicity.squarefree_iff_emultiplicity_le_one := squarefree_iff_emultiplicity_le_one section Irreducible @@ -266,7 +250,7 @@ lemma _root_.exists_squarefree_dvd_pow_of_ne_zero {x : R} (hx : x ≠ 0) : theorem squarefree_iff_nodup_normalizedFactors [NormalizationMonoid R] {x : R} (x0 : x ≠ 0) : Squarefree x ↔ Multiset.Nodup (normalizedFactors x) := by classical - rw [multiplicity.squarefree_iff_emultiplicity_le_one, Multiset.nodup_iff_count_le_one] + rw [squarefree_iff_emultiplicity_le_one, Multiset.nodup_iff_count_le_one] haveI := nontrivial_of_ne x 0 x0 constructor <;> intro h a · by_cases hmem : a ∈ normalizedFactors x diff --git a/Mathlib/Data/Nat/Choose/Factorization.lean b/Mathlib/Data/Nat/Choose/Factorization.lean index 6bd3a1dac00f3..e639e46d275bf 100644 --- a/Mathlib/Data/Nat/Choose/Factorization.lean +++ b/Mathlib/Data/Nat/Choose/Factorization.lean @@ -39,10 +39,10 @@ theorem factorization_choose_le_log : (choose n k).factorization p ≤ log p n : refine le_of_not_lt fun hnk => h ?_ simp [choose_eq_zero_of_lt hnk] rw [factorization_def _ hp, @padicValNat_def _ ⟨hp⟩ _ (choose_pos hkn)] - rw [← Nat.cast_le (α := ℕ∞), ← multiplicity.Finite.emultiplicity_eq_multiplicity] + rw [← Nat.cast_le (α := ℕ∞), ← FiniteMultiplicity.emultiplicity_eq_multiplicity] · simp only [hp.emultiplicity_choose hkn (lt_add_one _), Nat.cast_le] exact (Finset.card_filter_le _ _).trans (le_of_eq (Nat.card_Ico _ _)) - apply Nat.multiplicity_finite_iff.2 ⟨hp.ne_one, choose_pos hkn⟩ + apply Nat.finiteMultiplicity_iff.2 ⟨hp.ne_one, choose_pos hkn⟩ /-- A `pow` form of `Nat.factorization_choose_le` -/ theorem pow_factorization_choose_le (hn : 0 < n) : p ^ (choose n k).factorization p ≤ n := diff --git a/Mathlib/Data/Nat/Factorization/Basic.lean b/Mathlib/Data/Nat/Factorization/Basic.lean index 15659840ed9b9..ca5a516231dff 100644 --- a/Mathlib/Data/Nat/Factorization/Basic.lean +++ b/Mathlib/Data/Nat/Factorization/Basic.lean @@ -270,7 +270,7 @@ and `n'` such that `n'` is not divisible by `p` and `n = p^e * n'`. -/ theorem exists_eq_pow_mul_and_not_dvd {n : ℕ} (hn : n ≠ 0) (p : ℕ) (hp : p ≠ 1) : ∃ e n' : ℕ, ¬p ∣ n' ∧ n = p ^ e * n' := let ⟨a', h₁, h₂⟩ := - (Nat.multiplicity_finite_iff.mpr ⟨hp, Nat.pos_of_ne_zero hn⟩).exists_eq_pow_mul_and_not_dvd + (Nat.finiteMultiplicity_iff.mpr ⟨hp, Nat.pos_of_ne_zero hn⟩).exists_eq_pow_mul_and_not_dvd ⟨_, a', h₂, h₁⟩ /-- Any nonzero natural number is the product of an odd part `m` and a power of diff --git a/Mathlib/Data/Nat/Multiplicity.lean b/Mathlib/Data/Nat/Multiplicity.lean index 77e18051deac1..ece38a314c41c 100644 --- a/Mathlib/Data/Nat/Multiplicity.lean +++ b/Mathlib/Data/Nat/Multiplicity.lean @@ -46,7 +46,7 @@ Legendre, p-adic -/ -open Finset Nat multiplicity +open Finset Nat open Nat @@ -57,7 +57,7 @@ divides `n`. This set is expressed by filtering `Ico 1 b` where `b` is any bound `log m n`. -/ theorem emultiplicity_eq_card_pow_dvd {m n b : ℕ} (hm : m ≠ 1) (hn : 0 < n) (hb : log m n < b) : emultiplicity m n = #{i ∈ Ico 1 b | m ^ i ∣ n} := - have fin := Nat.multiplicity_finite_iff.2 ⟨hm, hn⟩ + have fin := Nat.finiteMultiplicity_iff.2 ⟨hm, hn⟩ calc emultiplicity m n = #(Ico 1 <| multiplicity m n + 1) := by simp [fin.emultiplicity_eq_multiplicity] @@ -91,7 +91,7 @@ theorem emultiplicity_pow {p m n : ℕ} (hp : p.Prime) : _root_.emultiplicity_pow hp.prime theorem emultiplicity_self {p : ℕ} (hp : p.Prime) : emultiplicity p p = 1 := - (Nat.multiplicity_finite_iff.2 ⟨hp.ne_one, hp.pos⟩).emultiplicity_self + (Nat.finiteMultiplicity_iff.2 ⟨hp.ne_one, hp.pos⟩).emultiplicity_self theorem emultiplicity_pow_self {p n : ℕ} (hp : p.Prime) : emultiplicity p (p ^ n) = n := _root_.emultiplicity_pow_self hp.ne_zero hp.prime.not_unit n @@ -136,7 +136,7 @@ theorem emultiplicity_factorial_mul_succ {n p : ℕ} (hp : p.Prime) : have h2 : p * n + 1 ≤ p * (n + 1) := by linarith have h3 : p * n + 1 ≤ p * (n + 1) + 1 := by omega have hm : emultiplicity p (p * n)! ≠ ⊤ := by - rw [Ne, emultiplicity_eq_top, Classical.not_not, Nat.multiplicity_finite_iff] + rw [Ne, emultiplicity_eq_top, Classical.not_not, Nat.finiteMultiplicity_iff] exact ⟨hp.ne_one, factorial_pos _⟩ revert hm have h4 : ∀ m ∈ Ico (p * n + 1) (p * (n + 1)), emultiplicity p m = 0 := by @@ -201,8 +201,8 @@ theorem emultiplicity_choose' {p n k b : ℕ} (hp : p.Prime) (hnb : log p (n + k (add_comm n k ▸ hnb)), multiplicity_choose_aux hp (le_add_left k n)] simp [add_comm] refine (WithTop.add_right_cancel_iff ?_).1 h₁ - apply finite_iff_emultiplicity_ne_top.1 - exact Nat.multiplicity_finite_iff.2 ⟨hp.ne_one, mul_pos (factorial_pos k) (factorial_pos n)⟩ + apply finiteMultiplicity_iff_emultiplicity_ne_top.1 + exact Nat.finiteMultiplicity_iff.2 ⟨hp.ne_one, mul_pos (factorial_pos k) (factorial_pos n)⟩ /-- The multiplicity of `p` in `choose n k` is the number of carries when `k` and `n - k` are added in base `p`. The set is expressed by filtering `Ico 1 b` where `b` @@ -251,8 +251,8 @@ theorem emultiplicity_choose_prime_pow {p n k : ℕ} (hp : p.Prime) (hkn : k ≤ emultiplicity p (choose (p ^ n) k) = ↑(n - multiplicity p k) := by push_cast rw [← emultiplicity_choose_prime_pow_add_emultiplicity hp hkn hk0, - (multiplicity_finite_iff.2 ⟨hp.ne_one, Nat.pos_of_ne_zero hk0⟩).emultiplicity_eq_multiplicity, - (multiplicity_finite_iff.2 ⟨hp.ne_one, choose_pos hkn⟩).emultiplicity_eq_multiplicity] + (finiteMultiplicity_iff.2 ⟨hp.ne_one, Nat.pos_of_ne_zero hk0⟩).emultiplicity_eq_multiplicity, + (finiteMultiplicity_iff.2 ⟨hp.ne_one, choose_pos hkn⟩).emultiplicity_eq_multiplicity] norm_cast rw [Nat.add_sub_cancel_right] diff --git a/Mathlib/Data/Nat/Squarefree.lean b/Mathlib/Data/Nat/Squarefree.lean index 53bb8bb29fec5..35e092c930493 100644 --- a/Mathlib/Data/Nat/Squarefree.lean +++ b/Mathlib/Data/Nat/Squarefree.lean @@ -50,7 +50,7 @@ theorem _root_.Squarefree.natFactorization_le_one {n : ℕ} (p : ℕ) (hn : Squa n.factorization p ≤ 1 := by rcases eq_or_ne n 0 with (rfl | hn') · simp - rw [multiplicity.squarefree_iff_emultiplicity_le_one] at hn + rw [squarefree_iff_emultiplicity_le_one] at hn by_cases hp : p.Prime · have := hn p rw [← multiplicity_eq_factorization hp hn'] diff --git a/Mathlib/Data/Real/Irrational.lean b/Mathlib/Data/Real/Irrational.lean index 235dc1f4d78d5..12f99f0fcb687 100644 --- a/Mathlib/Data/Real/Irrational.lean +++ b/Mathlib/Data/Real/Irrational.lean @@ -25,7 +25,7 @@ but this only works if you `unseal Nat.sqrt.iter in` before the theorem where yo -/ -open Rat Real multiplicity +open Rat Real /-- A real number is irrational if it is not equal to any rational number. -/ def Irrational (x : ℝ) := @@ -79,7 +79,7 @@ theorem irrational_nrt_of_n_not_dvd_multiplicity {x : ℝ} (n : ℕ) {m : ℤ} ( rw [← Int.cast_pow, Int.cast_inj] at hxr subst m have : y ≠ 0 := by rintro rfl; rw [zero_pow hnpos.ne'] at hm; exact hm rfl - rw [(Int.multiplicity_finite_iff.2 ⟨by simp [hp.1.ne_one], this⟩).multiplicity_pow + rw [(Int.finiteMultiplicity_iff.2 ⟨by simp [hp.1.ne_one], this⟩).multiplicity_pow (Nat.prime_iff_prime_int.1 hp.1), Nat.mul_mod_right] at hv exact hv rfl diff --git a/Mathlib/FieldTheory/Separable.lean b/Mathlib/FieldTheory/Separable.lean index 7488ef0514690..39db5e7df739e 100644 --- a/Mathlib/FieldTheory/Separable.lean +++ b/Mathlib/FieldTheory/Separable.lean @@ -184,7 +184,7 @@ See `PerfectField.separable_iff_squarefree` for the converse when the coefficien field. -/ theorem Separable.squarefree {p : R[X]} (hsep : Separable p) : Squarefree p := by classical - rw [multiplicity.squarefree_iff_emultiplicity_le_one p] + rw [squarefree_iff_emultiplicity_le_one p] exact fun f => or_iff_not_imp_right.mpr fun hunit => emultiplicity_le_one_of_separable hunit hsep end CommSemiring @@ -282,7 +282,7 @@ theorem rootMultiplicity_le_one_of_separable [Nontrivial R] {p : R[X]} (hsep : S by_cases hp : p = 0 · simp [hp] rw [rootMultiplicity_eq_multiplicity, if_neg hp, ← Nat.cast_le (α := ℕ∞), - Nat.cast_one, ← (multiplicity_X_sub_C_finite x hp).emultiplicity_eq_multiplicity] + Nat.cast_one, ← (finiteMultiplicity_X_sub_C x hp).emultiplicity_eq_multiplicity] apply emultiplicity_le_one_of_separable (not_isUnit_X_sub_C _) hsep end CommRing diff --git a/Mathlib/FieldTheory/SeparableDegree.lean b/Mathlib/FieldTheory/SeparableDegree.lean index 1b85e1f399fbf..8cc00db3d0c2b 100644 --- a/Mathlib/FieldTheory/SeparableDegree.lean +++ b/Mathlib/FieldTheory/SeparableDegree.lean @@ -592,7 +592,7 @@ theorem eq_X_pow_char_pow_sub_C_pow_of_natSepDegree_eq_one (q : ℕ) [ExpChar F have hD := (h ▸ natSepDegree_le_of_dvd p f hf hm.ne_zero).antisymm <| Nat.pos_of_ne_zero <| (natSepDegree_ne_zero_iff _).2 hI.natDegree_pos.ne' obtain ⟨n, y, H, hp⟩ := hM.eq_X_pow_char_pow_sub_C_of_natSepDegree_eq_one_of_irreducible q hI hD - have hF := multiplicity_finite_of_degree_pos_of_monic (degree_pos_of_irreducible hI) hM hm.ne_zero + have hF := finiteMultiplicity_of_degree_pos_of_monic (degree_pos_of_irreducible hI) hM hm.ne_zero classical have hne := (multiplicity_pos_of_dvd hf).ne' refine ⟨_, n, y, hne, H, ?_⟩ diff --git a/Mathlib/NumberTheory/FLT/Three.lean b/Mathlib/NumberTheory/FLT/Three.lean index 5ec4e154d21a3..fc457ccd599f2 100644 --- a/Mathlib/NumberTheory/FLT/Three.lean +++ b/Mathlib/NumberTheory/FLT/Three.lean @@ -210,8 +210,8 @@ variable [NumberField K] [IsCyclotomicExtension {3} ℚ K] /-- For any `S' : Solution'`, the multiplicity of `λ` in `S'.c` is finite. -/ lemma Solution'.multiplicity_lambda_c_finite : - multiplicity.Finite (hζ.toInteger - 1) S'.c := - multiplicity.finite_of_not_isUnit hζ.zeta_sub_one_prime'.not_unit S'.hc + FiniteMultiplicity (hζ.toInteger - 1) S'.c := + .of_not_isUnit hζ.zeta_sub_one_prime'.not_unit S'.hc section DecidableRel @@ -328,14 +328,14 @@ lemma lambda_sq_dvd_or_dvd_or_dvd : by_contra! h rcases h with ⟨h1, h2, h3⟩ rw [← emultiplicity_lt_iff_not_dvd] at h1 h2 h3 - have h1' : multiplicity.Finite (hζ.toInteger - 1) (S'.a + S'.b) := - finite_iff_emultiplicity_ne_top.2 (fun ht ↦ by simp [ht] at h1) - have h2' : multiplicity.Finite (hζ.toInteger - 1) (S'.a + η * S'.b) := by - refine finite_iff_emultiplicity_ne_top.2 (fun ht ↦ ?_) + have h1' : FiniteMultiplicity (hζ.toInteger - 1) (S'.a + S'.b) := + finiteMultiplicity_iff_emultiplicity_ne_top.2 (fun ht ↦ by simp [ht] at h1) + have h2' : FiniteMultiplicity (hζ.toInteger - 1) (S'.a + η * S'.b) := by + refine finiteMultiplicity_iff_emultiplicity_ne_top.2 (fun ht ↦ ?_) rw [coe_eta] at ht simp [ht] at h2 - have h3' : multiplicity.Finite (hζ.toInteger - 1) (S'.a + η ^ 2 * S'.b) := by - refine finite_iff_emultiplicity_ne_top.2 (fun ht ↦ ?_) + have h3' : FiniteMultiplicity (hζ.toInteger - 1) (S'.a + η ^ 2 * S'.b) := by + refine finiteMultiplicity_iff_emultiplicity_ne_top.2 (fun ht ↦ ?_) rw [coe_eta] at ht simp [ht] at h3 rw [h1'.emultiplicity_eq_multiplicity, Nat.cast_lt] at h1 diff --git a/Mathlib/NumberTheory/Multiplicity.lean b/Mathlib/NumberTheory/Multiplicity.lean index 85a6e500e4a62..ab751110fdab3 100644 --- a/Mathlib/NumberTheory/Multiplicity.lean +++ b/Mathlib/NumberTheory/Multiplicity.lean @@ -141,21 +141,24 @@ theorem odd_sq_dvd_geom_sum₂_sub (hp : Odd p) : refine Ideal.Quotient.eq_zero_iff_mem.mpr ?_ simp [mem_span_singleton] -namespace multiplicity - section IntegralDomain variable [IsDomain R] -theorem pow_sub_pow_of_prime {p : R} (hp : Prime p) {x y : R} (hxy : p ∣ x - y) (hx : ¬p ∣ x) - {n : ℕ} (hn : ¬p ∣ n) : emultiplicity p (x ^ n - y ^ n) = emultiplicity p (x - y) := by +theorem emultiplicity_pow_sub_pow_of_prime {p : R} (hp : Prime p) {x y : R} + (hxy : p ∣ x - y) (hx : ¬p ∣ x) {n : ℕ} (hn : ¬p ∣ n) : + emultiplicity p (x ^ n - y ^ n) = emultiplicity p (x - y) := by rw [← geom_sum₂_mul, emultiplicity_mul hp, emultiplicity_eq_zero.2 (not_dvd_geom_sum₂ hp hxy hx hn), zero_add] +@[deprecated (since := "2024-11-30")] +alias multiplicity.pow_sub_pow_of_prime := emultiplicity_pow_sub_pow_of_prime + variable (hp : Prime (p : R)) (hp1 : Odd p) (hxy : ↑p ∣ x - y) (hx : ¬↑p ∣ x) include hp hp1 hxy hx -theorem geom_sum₂_eq_one : emultiplicity (↑p) (∑ i ∈ range p, x ^ i * y ^ (p - 1 - i)) = 1 := by +theorem emultiplicity_geom_sum₂_eq_one : + emultiplicity (↑p) (∑ i ∈ range p, x ^ i * y ^ (p - 1 - i)) = 1 := by rw [← Nat.cast_one] refine emultiplicity_eq_coe.2 ⟨?_, ?_⟩ · rw [pow_one] @@ -167,20 +170,29 @@ theorem geom_sum₂_eq_one : emultiplicity (↑p) (∑ i ∈ range p, x ^ i * y rw [pow_two, mul_dvd_mul_iff_left hp.ne_zero] exact mt hp.dvd_of_dvd_pow hx -theorem pow_prime_sub_pow_prime : +@[deprecated (since := "2024-11-30")] +alias multiplicity.geom_sum₂_eq_one := emultiplicity_geom_sum₂_eq_one + +theorem emultiplicity_pow_prime_sub_pow_prime : emultiplicity (↑p) (x ^ p - y ^ p) = emultiplicity (↑p) (x - y) + 1 := by - rw [← geom_sum₂_mul, emultiplicity_mul hp, geom_sum₂_eq_one hp hp1 hxy hx, add_comm] + rw [← geom_sum₂_mul, emultiplicity_mul hp, emultiplicity_geom_sum₂_eq_one hp hp1 hxy hx, add_comm] -theorem pow_prime_pow_sub_pow_prime_pow (a : ℕ) : +@[deprecated (since := "2024-11-30")] +alias multiplicity.pow_prime_sub_pow_prime := emultiplicity_pow_prime_sub_pow_prime + +theorem emultiplicity_pow_prime_pow_sub_pow_prime_pow (a : ℕ) : emultiplicity (↑p) (x ^ p ^ a - y ^ p ^ a) = emultiplicity (↑p) (x - y) + a := by induction' a with a h_ind · rw [Nat.cast_zero, add_zero, pow_zero, pow_one, pow_one] rw [Nat.cast_add, Nat.cast_one, ← add_assoc, ← h_ind, pow_succ, pow_mul, pow_mul] - apply pow_prime_sub_pow_prime hp hp1 + apply emultiplicity_pow_prime_sub_pow_prime hp hp1 · rw [← geom_sum₂_mul] exact dvd_mul_of_dvd_right hxy _ · exact fun h => hx (hp.dvd_of_dvd_pow h) +@[deprecated (since := "2024-11-30")] +alias multiplicity.pow_prime_pow_sub_pow_prime_pow := emultiplicity_pow_prime_pow_sub_pow_prime_pow + end IntegralDomain section LiftingTheExponent @@ -189,17 +201,17 @@ variable (hp : Nat.Prime p) (hp1 : Odd p) include hp hp1 /-- **Lifting the exponent lemma** for odd primes. -/ -theorem Int.pow_sub_pow {x y : ℤ} (hxy : ↑p ∣ x - y) (hx : ¬↑p ∣ x) (n : ℕ) : +theorem Int.emultiplicity_pow_sub_pow {x y : ℤ} (hxy : ↑p ∣ x - y) (hx : ¬↑p ∣ x) (n : ℕ) : emultiplicity (↑p) (x ^ n - y ^ n) = emultiplicity (↑p) (x - y) + emultiplicity p n := by cases' n with n · simp only [emultiplicity_zero, add_top, pow_zero, sub_self] - have h : Finite _ _ := Nat.multiplicity_finite_iff.mpr ⟨hp.ne_one, n.succ_pos⟩ + have h : FiniteMultiplicity _ _ := Nat.finiteMultiplicity_iff.mpr ⟨hp.ne_one, n.succ_pos⟩ simp only [Nat.succ_eq_add_one] at h rcases emultiplicity_eq_coe.mp h.emultiplicity_eq_multiplicity with ⟨⟨k, hk⟩, hpn⟩ conv_lhs => rw [hk, pow_mul, pow_mul] rw [Nat.prime_iff_prime_int] at hp - rw [pow_sub_pow_of_prime hp, pow_prime_pow_sub_pow_prime_pow hp hp1 hxy hx, - h.emultiplicity_eq_multiplicity] + rw [emultiplicity_pow_sub_pow_of_prime hp, + emultiplicity_pow_prime_pow_sub_pow_prime_pow hp hp1 hxy hx, h.emultiplicity_eq_multiplicity] · rw [← geom_sum₂_mul] exact dvd_mul_of_dvd_right hxy _ · exact fun h => hx (hp.dvd_of_dvd_pow h) @@ -208,13 +220,20 @@ theorem Int.pow_sub_pow {x y : ℤ} (hxy : ↑p ∣ x - y) (hx : ¬↑p ∣ x) ( refine hpn ⟨c, ?_⟩ rwa [pow_succ, mul_assoc] -theorem Int.pow_add_pow {x y : ℤ} (hxy : ↑p ∣ x + y) (hx : ¬↑p ∣ x) {n : ℕ} (hn : Odd n) : +@[deprecated (since := "2024-11-30")] +alias multiplicity.Int.pow_sub_pow := Int.emultiplicity_pow_sub_pow + +theorem Int.emultiplicity_pow_add_pow {x y : ℤ} (hxy : ↑p ∣ x + y) (hx : ¬↑p ∣ x) + {n : ℕ} (hn : Odd n) : emultiplicity (↑p) (x ^ n + y ^ n) = emultiplicity (↑p) (x + y) + emultiplicity p n := by rw [← sub_neg_eq_add] at hxy rw [← sub_neg_eq_add, ← sub_neg_eq_add, ← Odd.neg_pow hn] - exact Int.pow_sub_pow hp hp1 hxy hx n + exact Int.emultiplicity_pow_sub_pow hp hp1 hxy hx n + +@[deprecated (since := "2024-11-30")] +alias multiplicity.Int.pow_add_pow := Int.emultiplicity_pow_add_pow -theorem Nat.pow_sub_pow {x y : ℕ} (hxy : p ∣ x - y) (hx : ¬p ∣ x) (n : ℕ) : +theorem Nat.emultiplicity_pow_sub_pow {x y : ℕ} (hxy : p ∣ x - y) (hx : ¬p ∣ x) (n : ℕ) : emultiplicity p (x ^ n - y ^ n) = emultiplicity p (x - y) + emultiplicity p n := by obtain hyx | hyx := le_total y x · iterate 2 rw [← Int.natCast_emultiplicity] @@ -222,20 +241,25 @@ theorem Nat.pow_sub_pow {x y : ℕ} (hxy : p ∣ x - y) (hx : ¬p ∣ x) (n : rw [← Int.natCast_dvd_natCast] at hxy hx rw [Int.natCast_sub hyx] at * push_cast at * - exact Int.pow_sub_pow hp hp1 hxy hx n + exact Int.emultiplicity_pow_sub_pow hp hp1 hxy hx n · simp only [Nat.sub_eq_zero_iff_le.mpr (Nat.pow_le_pow_left hyx n), emultiplicity_zero, Nat.sub_eq_zero_iff_le.mpr hyx, top_add] -theorem Nat.pow_add_pow {x y : ℕ} (hxy : p ∣ x + y) (hx : ¬p ∣ x) {n : ℕ} (hn : Odd n) : +@[deprecated (since := "2024-11-30")] +alias multiplicity.Nat.pow_sub_pow := Nat.emultiplicity_pow_sub_pow + +theorem Nat.emultiplicity_pow_add_pow {x y : ℕ} (hxy : p ∣ x + y) (hx : ¬p ∣ x) + {n : ℕ} (hn : Odd n) : emultiplicity p (x ^ n + y ^ n) = emultiplicity p (x + y) + emultiplicity p n := by iterate 2 rw [← Int.natCast_emultiplicity] rw [← Int.natCast_dvd_natCast] at hxy hx push_cast at * - exact Int.pow_add_pow hp hp1 hxy hx hn + exact Int.emultiplicity_pow_add_pow hp hp1 hxy hx hn -end LiftingTheExponent +@[deprecated (since := "2024-11-30")] +alias multiplicity.Nat.pow_add_pow := Nat.emultiplicity_pow_add_pow -end multiplicity +end LiftingTheExponent end CommRing @@ -294,10 +318,10 @@ theorem Int.two_pow_sub_pow' {x y : ℤ} (n : ℕ) (hxy : 4 ∣ x - y) (hx : ¬2 have hy_odd : Odd y := by simpa using hx_odd.sub_even hxy_even cases' n with n · simp only [pow_zero, sub_self, emultiplicity_zero, Int.ofNat_zero, add_top] - have h : multiplicity.Finite 2 n.succ := Nat.multiplicity_finite_iff.mpr ⟨by norm_num, n.succ_pos⟩ + have h : FiniteMultiplicity 2 n.succ := Nat.finiteMultiplicity_iff.mpr ⟨by norm_num, n.succ_pos⟩ simp only [Nat.succ_eq_add_one] at h rcases emultiplicity_eq_coe.mp h.emultiplicity_eq_multiplicity with ⟨⟨k, hk⟩, hpn⟩ - rw [hk, pow_mul, pow_mul, multiplicity.pow_sub_pow_of_prime, + rw [hk, pow_mul, pow_mul, emultiplicity_pow_sub_pow_of_prime, Int.two_pow_two_pow_sub_pow_two_pow _ hxy hx, ← hk] · norm_cast rw [h.emultiplicity_eq_multiplicity] @@ -332,8 +356,8 @@ theorem Int.two_pow_sub_pow {x y : ℤ} {n : ℕ} (hxy : 2 ∣ x - y) (hx : ¬2 emultiplicity_mul Int.prime_two, emultiplicity_mul Int.prime_two] · suffices emultiplicity (2 : ℤ) ↑(2 : ℕ) = 1 by rw [this, add_comm 1, ← add_assoc] norm_cast - rw [multiplicity.Finite.emultiplicity_self] - rw [Nat.multiplicity_finite_iff] + rw [FiniteMultiplicity.emultiplicity_self] + rw [Nat.finiteMultiplicity_iff] decide · rw [← even_iff_two_dvd, Int.not_even_iff_odd] apply Odd.pow @@ -378,7 +402,7 @@ theorem pow_sub_pow (hyx : y < x) (hxy : p ∣ x - y) (hx : ¬p ∣ x) {n : ℕ} padicValNat p (x ^ n - y ^ n) = padicValNat p (x - y) + padicValNat p n := by rw [← Nat.cast_inj (R := ℕ∞), Nat.cast_add] iterate 3 rw [padicValNat_eq_emultiplicity] - · exact multiplicity.Nat.pow_sub_pow hp.out hp1 hxy hx n + · exact Nat.emultiplicity_pow_sub_pow hp.out hp1 hxy hx n · exact hn.bot_lt · exact Nat.sub_pos_of_lt hyx · exact Nat.sub_pos_of_lt (Nat.pow_lt_pow_left hyx hn) @@ -389,7 +413,7 @@ theorem pow_add_pow (hxy : p ∣ x + y) (hx : ¬p ∣ x) {n : ℕ} (hn : Odd n) · contradiction rw [← Nat.cast_inj (R := ℕ∞), Nat.cast_add] iterate 3 rw [padicValNat_eq_emultiplicity] - · exact multiplicity.Nat.pow_add_pow hp.out hp1 hxy hx hn + · exact Nat.emultiplicity_pow_add_pow hp.out hp1 hxy hx hn · exact Odd.pos hn · simp only [add_pos_iff, Nat.succ_pos', or_true] · exact Nat.lt_add_left _ (pow_pos y.succ_pos _) diff --git a/Mathlib/NumberTheory/Padics/PadicNorm.lean b/Mathlib/NumberTheory/Padics/PadicNorm.lean index a50570c41e9f1..5e8647a2a1598 100644 --- a/Mathlib/NumberTheory/Padics/PadicNorm.lean +++ b/Mathlib/NumberTheory/Padics/PadicNorm.lean @@ -227,9 +227,9 @@ theorem dvd_iff_norm_le {n : ℕ} {z : ℤ} : ↑(p ^ n) ∣ z ↔ padicNorm p z · rw [zpow_le_zpow_iff_right₀, neg_le_neg_iff, padicValRat.of_int, padicValInt.of_ne_one_ne_zero hp.1.ne_one _] · norm_cast - rw [← multiplicity.Finite.pow_dvd_iff_le_multiplicity] + rw [← FiniteMultiplicity.pow_dvd_iff_le_multiplicity] · norm_cast - · apply Int.multiplicity_finite_iff.2 ⟨by simp [hp.out.ne_one], mod_cast hz⟩ + · apply Int.finiteMultiplicity_iff.2 ⟨by simp [hp.out.ne_one], mod_cast hz⟩ · exact_mod_cast hz · exact_mod_cast hp.out.one_lt diff --git a/Mathlib/NumberTheory/Padics/PadicNumbers.lean b/Mathlib/NumberTheory/Padics/PadicNumbers.lean index d86fe1b155206..7e461266cf61d 100644 --- a/Mathlib/NumberTheory/Padics/PadicNumbers.lean +++ b/Mathlib/NumberTheory/Padics/PadicNumbers.lean @@ -60,7 +60,7 @@ p-adic, p adic, padic, norm, valuation, cauchy, completion, p-adic completion noncomputable section -open Nat multiplicity padicNorm CauSeq CauSeq.Completion Metric +open Nat padicNorm CauSeq CauSeq.Completion Metric /-- The type of Cauchy sequences of rationals with respect to the `p`-adic norm. -/ abbrev PadicSeq (p : ℕ) := diff --git a/Mathlib/NumberTheory/Padics/PadicVal/Basic.lean b/Mathlib/NumberTheory/Padics/PadicVal/Basic.lean index d2df11023e6f1..07c80e89edce7 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal/Basic.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal/Basic.lean @@ -66,12 +66,8 @@ open Nat open Rat -open multiplicity - namespace padicValNat -open multiplicity - variable {p : ℕ} /-- If `p ≠ 0` and `p ≠ 1`, then `padicValNat p p` is `1`. -/ @@ -91,7 +87,7 @@ theorem maxPowDiv_eq_emultiplicity {p n : ℕ} (hp : 1 < p) (hn : 0 < n) : apply Nat.not_lt.mpr <| le_of_dvd hp hn h simp -theorem maxPowDiv_eq_multiplicity {p n : ℕ} (hp : 1 < p) (hn : 0 < n) (h : Finite p n) : +theorem maxPowDiv_eq_multiplicity {p n : ℕ} (hp : 1 < p) (hn : 0 < n) (h : FiniteMultiplicity p n) : p.maxPowDiv n = multiplicity p n := by exact_mod_cast h.emultiplicity_eq_multiplicity ▸ maxPowDiv_eq_emultiplicity hp hn @@ -101,7 +97,7 @@ theorem padicValNat_eq_maxPowDiv : @padicValNat = @maxPowDiv := by ext p n by_cases h : 1 < p ∧ 0 < n · rw [padicValNat_def' h.1.ne' h.2, maxPowDiv_eq_multiplicity h.1 h.2] - exact Nat.multiplicity_finite_iff.2 ⟨h.1.ne', h.2⟩ + exact Nat.finiteMultiplicity_iff.2 ⟨h.1.ne', h.2⟩ · simp only [not_and_or,not_gt_eq,Nat.le_zero] at h apply h.elim · intro h @@ -121,8 +117,6 @@ def padicValInt (p : ℕ) (z : ℤ) : ℕ := namespace padicValInt -open multiplicity - variable {p : ℕ} theorem of_ne_one_ne_zero {z : ℤ} (hp : p ≠ 1) (hz : z ≠ 0) : @@ -163,8 +157,6 @@ lemma padicValRat_def (p : ℕ) (q : ℚ) : namespace padicValRat -open multiplicity - variable {p : ℕ} /-- `padicValRat p q` is symmetric in `q`. -/ @@ -232,13 +224,11 @@ end padicValNat namespace padicValRat -open multiplicity - variable {p : ℕ} [hp : Fact p.Prime] /-- The multiplicity of `p : ℕ` in `a : ℤ` is finite exactly when `a ≠ 0`. -/ -theorem finite_int_prime_iff {a : ℤ} : Finite (p : ℤ) a ↔ a ≠ 0 := by - simp [Int.multiplicity_finite_iff, hp.1.ne_one] +theorem finite_int_prime_iff {a : ℤ} : FiniteMultiplicity (p : ℤ) a ↔ a ≠ 0 := by + simp [Int.finiteMultiplicity_iff, hp.1.ne_one] /-- A rewrite lemma for `padicValRat p q` when `q` is expressed in terms of `Rat.mk`. -/ protected theorem defn (p : ℕ) [hp : Fact p.Prime] {q : ℚ} {n d : ℤ} (hqz : q ≠ 0) @@ -296,8 +286,8 @@ theorem padicValRat_le_padicValRat_iff {n₁ n₂ d₁ d₂ : ℤ} (hn₁ : n₁ (hd₁ : d₁ ≠ 0) (hd₂ : d₂ ≠ 0) : padicValRat p (n₁ /. d₁) ≤ padicValRat p (n₂ /. d₂) ↔ ∀ n : ℕ, (p : ℤ) ^ n ∣ n₁ * d₂ → (p : ℤ) ^ n ∣ n₂ * d₁ := by - have hf1 : Finite (p : ℤ) (n₁ * d₂) := finite_int_prime_iff.2 (mul_ne_zero hn₁ hd₂) - have hf2 : Finite (p : ℤ) (n₂ * d₁) := finite_int_prime_iff.2 (mul_ne_zero hn₂ hd₁) + have hf1 : FiniteMultiplicity (p : ℤ) (n₁ * d₂) := finite_int_prime_iff.2 (mul_ne_zero hn₁ hd₂) + have hf2 : FiniteMultiplicity (p : ℤ) (n₂ * d₁) := finite_int_prime_iff.2 (mul_ne_zero hn₂ hd₁) conv => lhs rw [padicValRat.defn p (Rat.divInt_ne_zero_of_ne_zero hn₁ hd₁) rfl, diff --git a/Mathlib/NumberTheory/Padics/PadicVal/Defs.lean b/Mathlib/NumberTheory/Padics/PadicVal/Defs.lean index cba87c009c504..39a5bdc501f09 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal/Defs.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal/Defs.lean @@ -22,18 +22,16 @@ universe u open Nat -open multiplicity - variable {p : ℕ} /-- For `p ≠ 1`, the `p`-adic valuation of a natural `n ≠ 0` is the largest natural number `k` such that `p^k` divides `n`. If `n = 0` or `p = 1`, then `padicValNat p q` defaults to `0`. -/ def padicValNat (p : ℕ) (n : ℕ) : ℕ := - if h : p ≠ 1 ∧ 0 < n then Nat.find (multiplicity_finite_iff.2 h) else 0 + if h : p ≠ 1 ∧ 0 < n then Nat.find (finiteMultiplicity_iff.2 h) else 0 theorem padicValNat_def' {n : ℕ} (hp : p ≠ 1) (hn : 0 < n) : padicValNat p n = multiplicity p n := by - simp [padicValNat, hp, hn, multiplicity, emultiplicity, multiplicity_finite_iff.2 ⟨hp, hn⟩] + simp [padicValNat, hp, hn, multiplicity, emultiplicity, finiteMultiplicity_iff.2 ⟨hp, hn⟩] convert (WithTop.untop'_coe ..).symm /-- A simplification of `padicValNat` when one input is prime, by analogy with @@ -46,13 +44,11 @@ theorem padicValNat_def [hp : Fact p.Prime] {n : ℕ} (hn : 0 < n) : `padicValRat_def`. -/ theorem padicValNat_eq_emultiplicity [hp : Fact p.Prime] {n : ℕ} (hn : 0 < n) : padicValNat p n = emultiplicity p n := by - rw [(multiplicity_finite_iff.2 ⟨hp.out.ne_one, hn⟩).emultiplicity_eq_multiplicity] + rw [(finiteMultiplicity_iff.2 ⟨hp.out.ne_one, hn⟩).emultiplicity_eq_multiplicity] exact_mod_cast padicValNat_def hn namespace padicValNat -open multiplicity - open List /-- `padicValNat p 0` is `0` for any `p`. -/ @@ -86,7 +82,7 @@ theorem le_padicValNat_iff_replicate_subperm_primeFactorsList {a b : ℕ} {n : (hb : b ≠ 0) : n ≤ padicValNat a b ↔ replicate n a <+~ b.primeFactorsList := by rw [← le_emultiplicity_iff_replicate_subperm_primeFactorsList ha hb, - Nat.multiplicity_finite_iff.2 ⟨ha.ne_one, Nat.pos_of_ne_zero hb⟩ + Nat.finiteMultiplicity_iff.2 ⟨ha.ne_one, Nat.pos_of_ne_zero hb⟩ |>.emultiplicity_eq_multiplicity, ← padicValNat_def' ha.ne_one (Nat.pos_of_ne_zero hb), Nat.cast_le] diff --git a/Mathlib/RingTheory/ChainOfDivisors.lean b/Mathlib/RingTheory/ChainOfDivisors.lean index a1921d9752ce6..68af812321e63 100644 --- a/Mathlib/RingTheory/ChainOfDivisors.lean +++ b/Mathlib/RingTheory/ChainOfDivisors.lean @@ -5,7 +5,7 @@ Authors: Anne Baanen, Paul Lezeau -/ import Mathlib.Algebra.GCDMonoid.Basic import Mathlib.Algebra.IsPrimePow -import Mathlib.Algebra.Squarefree.Basic +import Mathlib.RingTheory.UniqueFactorizationDomain.Multiplicity import Mathlib.Data.ZMod.Defs import Mathlib.Order.Atoms import Mathlib.Order.Hom.Bounded @@ -58,7 +58,7 @@ theorem Associates.isAtom_iff {p : Associates M} (h₁ : p ≠ 0) : IsAtom p ↔ ⟨(ha.unit⁻¹ : Units _), by rw [hab, mul_assoc, IsUnit.mul_val_inv ha, mul_one]⟩) hb⟩⟩ -open UniqueFactorizationMonoid multiplicity Irreducible Associates +open UniqueFactorizationMonoid Irreducible Associates namespace DivisorChain @@ -311,8 +311,8 @@ theorem emultiplicity_prime_le_emultiplicity_image_by_factor_orderIso {m p : Ass · simp [hn] by_cases hm : m = 0 · simp [hm] at hp - rw [(finite_prime_left (prime_of_normalized_factor p hp) hm).emultiplicity_eq_multiplicity, - ← pow_dvd_iff_le_emultiplicity] + rw [FiniteMultiplicity.of_prime_left (prime_of_normalized_factor p hp) hm + |>.emultiplicity_eq_multiplicity, ← pow_dvd_iff_le_emultiplicity] apply pow_image_of_prime_by_factor_orderIso_dvd hn hp d (pow_multiplicity_dvd ..) theorem emultiplicity_prime_eq_emultiplicity_image_by_factor_orderIso {m p : Associates M} diff --git a/Mathlib/RingTheory/DedekindDomain/Ideal.lean b/Mathlib/RingTheory/DedekindDomain/Ideal.lean index 2533459ec4771..5925c3017eae3 100644 --- a/Mathlib/RingTheory/DedekindDomain/Ideal.lean +++ b/Mathlib/RingTheory/DedekindDomain/Ideal.lean @@ -911,7 +911,7 @@ theorem irreducible_pow_sup_of_ge (hI : I ≠ ⊥) (hJ : Irreducible J) (n : ℕ classical rw [irreducible_pow_sup hI hJ, min_eq_left] · congr - rw [← Nat.cast_inj (R := ℕ∞), ← multiplicity.Finite.emultiplicity_eq_multiplicity, + rw [← Nat.cast_inj (R := ℕ∞), ← FiniteMultiplicity.emultiplicity_eq_multiplicity, emultiplicity_eq_count_normalizedFactors hJ hI, normalize_eq J] rw [← emultiplicity_lt_top] apply hn.trans_lt @@ -1317,7 +1317,7 @@ end ChineseRemainder section PID -open multiplicity UniqueFactorizationMonoid Ideal +open UniqueFactorizationMonoid Ideal variable {R} variable [IsDomain R] [IsPrincipalIdealRing R] @@ -1358,18 +1358,18 @@ theorem singleton_span_mem_normalizedFactors_of_mem_normalizedFactors [Normaliza theorem emultiplicity_eq_emultiplicity_span {a b : R} : emultiplicity (Ideal.span {a}) (Ideal.span ({b} : Set R)) = emultiplicity a b := by - by_cases h : Finite a b + by_cases h : FiniteMultiplicity a b · rw [h.emultiplicity_eq_multiplicity] apply emultiplicity_eq_of_dvd_of_not_dvd <;> rw [Ideal.span_singleton_pow, span_singleton_dvd_span_singleton_iff_dvd] · exact pow_multiplicity_dvd a b · apply h.not_pow_dvd_of_multiplicity_lt apply lt_add_one - · suffices ¬Finite (Ideal.span ({a} : Set R)) (Ideal.span ({b} : Set R)) by + · suffices ¬FiniteMultiplicity (Ideal.span ({a} : Set R)) (Ideal.span ({b} : Set R)) by rw [emultiplicity_eq_top.2 h, emultiplicity_eq_top.2 this] - exact Finite.not_iff_forall.mpr fun n => by + exact FiniteMultiplicity.not_iff_forall.mpr fun n => by rw [Ideal.span_singleton_pow, span_singleton_dvd_span_singleton_iff_dvd] - exact Finite.not_iff_forall.mp h n + exact FiniteMultiplicity.not_iff_forall.mp h n section NormalizationMonoid variable [NormalizationMonoid R] diff --git a/Mathlib/RingTheory/Multiplicity.lean b/Mathlib/RingTheory/Multiplicity.lean index f6528327fbc78..c1dc9a1ff5b51 100644 --- a/Mathlib/RingTheory/Multiplicity.lean +++ b/Mathlib/RingTheory/Multiplicity.lean @@ -22,7 +22,7 @@ several basic results on it. `n`. * `multiplicity a b`: a `ℕ`-valued version of `multiplicity`, defaulting for `1` instead of `⊤`. The reason for using `1` as a default value instead of `0` is to have `multiplicity_eq_zero_iff`. -* `multiplicity.Finite a b`: a predicate denoting that the multiplicity of `a` in `b` is finite. +* `FiniteMultiplicity a b`: a predicate denoting that the multiplicity of `a` in `b` is finite. -/ @@ -31,47 +31,56 @@ variable {α β : Type*} open Nat /-- `multiplicity.Finite a b` indicates that the multiplicity of `a` in `b` is finite. -/ -abbrev multiplicity.Finite [Monoid α] (a b : α) : Prop := +abbrev FiniteMultiplicity [Monoid α] (a b : α) : Prop := ∃ n : ℕ, ¬a ^ (n + 1) ∣ b +@[deprecated (since := "2024-11-30")] alias multiplicity.Finite := FiniteMultiplicity + open scoped Classical in /-- `emultiplicity a b` returns the largest natural number `n` such that `a ^ n ∣ b`, as an `ℕ∞`. If `∀ n, a ^ n ∣ b` then it returns `⊤`. -/ noncomputable def emultiplicity [Monoid α] (a b : α) : ℕ∞ := - if h : multiplicity.Finite a b then Nat.find h else ⊤ + if h : FiniteMultiplicity a b then Nat.find h else ⊤ /-- A `ℕ`-valued version of `emultiplicity`, returning `1` instead of `⊤`. -/ noncomputable def multiplicity [Monoid α] (a b : α) : ℕ := (emultiplicity a b).untop' 1 -open multiplicity - section Monoid variable [Monoid α] [Monoid β] {a b : α} @[simp] theorem emultiplicity_eq_top : - emultiplicity a b = ⊤ ↔ ¬Finite a b := by + emultiplicity a b = ⊤ ↔ ¬FiniteMultiplicity a b := by simp [emultiplicity] -theorem emultiplicity_lt_top {a b : α} : emultiplicity a b < ⊤ ↔ Finite a b := by +theorem emultiplicity_lt_top {a b : α} : emultiplicity a b < ⊤ ↔ FiniteMultiplicity a b := by simp [lt_top_iff_ne_top, emultiplicity_eq_top] -theorem finite_iff_emultiplicity_ne_top : - Finite a b ↔ emultiplicity a b ≠ ⊤ := by simp +theorem finiteMultiplicity_iff_emultiplicity_ne_top : + FiniteMultiplicity a b ↔ emultiplicity a b ≠ ⊤ := by simp + +@[deprecated (since := "2024-11-30")] +alias finite_iff_emultiplicity_ne_top := finiteMultiplicity_iff_emultiplicity_ne_top -alias ⟨multiplicity.Finite.emultiplicity_ne_top, _⟩ := finite_iff_emultiplicity_ne_top +alias ⟨FiniteMultiplicity.emultiplicity_ne_top, _⟩ := finite_iff_emultiplicity_ne_top + +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.emultiplicity_ne_top := FiniteMultiplicity.emultiplicity_ne_top @[deprecated (since := "2024-11-08")] -alias Finite.emultiplicity_ne_top := multiplicity.Finite.emultiplicity_ne_top +alias Finite.emultiplicity_ne_top := FiniteMultiplicity.emultiplicity_ne_top -theorem finite_of_emultiplicity_eq_natCast {n : ℕ} (h : emultiplicity a b = n) : - Finite a b := by +theorem finiteMultiplicity_of_emultiplicity_eq_natCast {n : ℕ} (h : emultiplicity a b = n) : + FiniteMultiplicity a b := by by_contra! nh rw [← emultiplicity_eq_top, h] at nh trivial +@[deprecated (since := "2024-11-30")] +alias finite_of_emultiplicity_eq_natCast := finiteMultiplicity_of_emultiplicity_eq_natCast + theorem multiplicity_eq_of_emultiplicity_eq_some {n : ℕ} (h : emultiplicity a b = n) : multiplicity a b = n := by simp [multiplicity, h] @@ -81,16 +90,24 @@ theorem emultiplicity_ne_of_multiplicity_ne {n : ℕ} : multiplicity a b ≠ n → emultiplicity a b ≠ n := mt multiplicity_eq_of_emultiplicity_eq_some -theorem multiplicity.Finite.emultiplicity_eq_multiplicity (h : Finite a b) : +theorem FiniteMultiplicity.emultiplicity_eq_multiplicity (h : FiniteMultiplicity a b) : emultiplicity a b = multiplicity a b := by cases hm : emultiplicity a b · simp [h] at hm rw [multiplicity_eq_of_emultiplicity_eq_some hm] -theorem multiplicity.Finite.emultiplicity_eq_iff_multiplicity_eq {n : ℕ} (h : Finite a b) : - emultiplicity a b = n ↔ multiplicity a b = n := by +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.emultiplicity_eq_multiplicity := + FiniteMultiplicity.emultiplicity_eq_multiplicity + +theorem FiniteMultiplicity.emultiplicity_eq_iff_multiplicity_eq {n : ℕ} + (h : FiniteMultiplicity a b) : emultiplicity a b = n ↔ multiplicity a b = n := by simp [h.emultiplicity_eq_multiplicity] +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.emultiplicity_eq_iff_multiplicity_eq := + FiniteMultiplicity.emultiplicity_eq_iff_multiplicity_eq + theorem emultiplicity_eq_iff_multiplicity_eq_of_ne_one {n : ℕ} (h : n ≠ 1) : emultiplicity a b = n ↔ multiplicity a b = n := by constructor @@ -103,14 +120,18 @@ theorem emultiplicity_eq_zero_iff_multiplicity_eq_zero : emultiplicity_eq_iff_multiplicity_eq_of_ne_one zero_ne_one @[simp] -theorem multiplicity_eq_one_of_not_finite (h : ¬Finite a b) : +theorem multiplicity_eq_one_of_not_finiteMultiplicity (h : ¬FiniteMultiplicity a b) : multiplicity a b = 1 := by simp [multiplicity, emultiplicity_eq_top.2 h] +@[deprecated (since := "2024-11-30")] +alias multiplicity_eq_one_of_not_finite := + multiplicity_eq_one_of_not_finiteMultiplicity + @[simp] theorem multiplicity_le_emultiplicity : multiplicity a b ≤ emultiplicity a b := by - by_cases hf : Finite a b + by_cases hf : FiniteMultiplicity a b · simp [hf.emultiplicity_eq_multiplicity] · simp [hf, emultiplicity_eq_top.2] @@ -124,52 +145,73 @@ theorem multiplicity_le_of_emultiplicity_le {n : ℕ} (h : emultiplicity a b ≤ multiplicity a b ≤ n := by exact_mod_cast multiplicity_le_emultiplicity.trans h -theorem multiplicity.Finite.emultiplicity_le_of_multiplicity_le (hfin : Finite a b) +theorem FiniteMultiplicity.emultiplicity_le_of_multiplicity_le (hfin : FiniteMultiplicity a b) {n : ℕ} (h : multiplicity a b ≤ n) : emultiplicity a b ≤ n := by rw [emultiplicity_eq_multiplicity hfin] assumption_mod_cast +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.emultiplicity_le_of_multiplicity_le := + FiniteMultiplicity.emultiplicity_le_of_multiplicity_le + theorem le_emultiplicity_of_le_multiplicity {n : ℕ} (h : n ≤ multiplicity a b) : n ≤ emultiplicity a b := by exact_mod_cast (WithTop.coe_mono h).trans multiplicity_le_emultiplicity -theorem multiplicity.Finite.le_multiplicity_of_le_emultiplicity (hfin : Finite a b) +theorem FiniteMultiplicity.le_multiplicity_of_le_emultiplicity (hfin : FiniteMultiplicity a b) {n : ℕ} (h : n ≤ emultiplicity a b) : n ≤ multiplicity a b := by rw [emultiplicity_eq_multiplicity hfin] at h assumption_mod_cast +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.le_multiplicity_of_le_emultiplicity := + FiniteMultiplicity.le_multiplicity_of_le_emultiplicity + theorem multiplicity_lt_of_emultiplicity_lt {n : ℕ} (h : emultiplicity a b < n) : multiplicity a b < n := by exact_mod_cast multiplicity_le_emultiplicity.trans_lt h -theorem multiplicity.Finite.emultiplicity_lt_of_multiplicity_lt (hfin : Finite a b) +theorem FiniteMultiplicity.emultiplicity_lt_of_multiplicity_lt (hfin : FiniteMultiplicity a b) {n : ℕ} (h : multiplicity a b < n) : emultiplicity a b < n := by rw [emultiplicity_eq_multiplicity hfin] assumption_mod_cast +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.emultiplicity_lt_of_multiplicity_lt := + FiniteMultiplicity.emultiplicity_lt_of_multiplicity_lt + theorem lt_emultiplicity_of_lt_multiplicity {n : ℕ} (h : n < multiplicity a b) : n < emultiplicity a b := by exact_mod_cast (WithTop.coe_strictMono h).trans_le multiplicity_le_emultiplicity -theorem multiplicity.Finite.lt_multiplicity_of_lt_emultiplicity (hfin : Finite a b) +theorem FiniteMultiplicity.lt_multiplicity_of_lt_emultiplicity (hfin : FiniteMultiplicity a b) {n : ℕ} (h : n < emultiplicity a b) : n < multiplicity a b := by rw [emultiplicity_eq_multiplicity hfin] at h assumption_mod_cast +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.lt_multiplicity_of_lt_emultiplicity := + FiniteMultiplicity.lt_multiplicity_of_lt_emultiplicity + theorem emultiplicity_pos_iff : 0 < emultiplicity a b ↔ 0 < multiplicity a b := by simp [pos_iff_ne_zero, pos_iff_ne_zero, emultiplicity_eq_zero_iff_multiplicity_eq_zero] -theorem multiplicity.Finite.def : Finite a b ↔ ∃ n : ℕ, ¬a ^ (n + 1) ∣ b := +theorem FiniteMultiplicity.def : FiniteMultiplicity a b ↔ ∃ n : ℕ, ¬a ^ (n + 1) ∣ b := Iff.rfl -theorem multiplicity.Finite.not_dvd_of_one_right : Finite a 1 → ¬a ∣ 1 := fun ⟨n, hn⟩ ⟨d, hd⟩ => - hn ⟨d ^ (n + 1), (pow_mul_pow_eq_one (n + 1) hd.symm).symm⟩ +@[deprecated (since := "2024-11-30")] alias multiplicity.Finite.def := FiniteMultiplicity.def + +theorem FiniteMultiplicity.not_dvd_of_one_right : FiniteMultiplicity a 1 → ¬a ∣ 1 := + fun ⟨n, hn⟩ ⟨d, hd⟩ => hn ⟨d ^ (n + 1), (pow_mul_pow_eq_one (n + 1) hd.symm).symm⟩ + +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.not_dvd_of_one_right := FiniteMultiplicity.not_dvd_of_one_right @[norm_cast] theorem Int.natCast_emultiplicity (a b : ℕ) : emultiplicity (a : ℤ) (b : ℤ) = emultiplicity a b := by - unfold emultiplicity multiplicity.Finite + unfold emultiplicity FiniteMultiplicity congr! <;> norm_cast @[norm_cast] @@ -178,22 +220,32 @@ theorem Int.natCast_multiplicity (a b : ℕ) : multiplicity (a : ℤ) (b : ℤ) @[deprecated (since := "2024-04-05")] alias Int.coe_nat_multiplicity := Int.natCast_multiplicity -theorem multiplicity.Finite.not_iff_forall : ¬Finite a b ↔ ∀ n : ℕ, a ^ n ∣ b := +theorem FiniteMultiplicity.not_iff_forall : ¬FiniteMultiplicity a b ↔ ∀ n : ℕ, a ^ n ∣ b := ⟨fun h n => Nat.casesOn n (by rw [_root_.pow_zero] exact one_dvd _) - (by simpa [multiplicity.Finite, Classical.not_not] using h), - by simp [multiplicity.Finite, multiplicity, Classical.not_not]; tauto⟩ + (by simpa [FiniteMultiplicity] using h), + by simp [FiniteMultiplicity, multiplicity]; tauto⟩ -theorem multiplicity.Finite.not_unit (h : Finite a b) : ¬IsUnit a := +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.not_iff_forall := FiniteMultiplicity.not_iff_forall + +theorem FiniteMultiplicity.not_unit (h : FiniteMultiplicity a b) : ¬IsUnit a := let ⟨n, hn⟩ := h hn ∘ IsUnit.dvd ∘ IsUnit.pow (n + 1) -theorem multiplicity.Finite.mul_left {c : α} : Finite a (b * c) → Finite a b := fun ⟨n, hn⟩ => +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.not_unit := FiniteMultiplicity.not_unit + +theorem FiniteMultiplicity.mul_left {c : α} : + FiniteMultiplicity a (b * c) → FiniteMultiplicity a b := fun ⟨n, hn⟩ => ⟨n, fun h => hn (h.trans (dvd_mul_right _ _))⟩ +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.mul_left := FiniteMultiplicity.mul_left + theorem pow_dvd_of_le_emultiplicity {k : ℕ} (hk : k ≤ emultiplicity a b) : a ^ k ∣ b := by classical cases k @@ -202,7 +254,7 @@ theorem pow_dvd_of_le_emultiplicity {k : ℕ} (hk : k ≤ emultiplicity a b) : split at hk · norm_cast at hk simpa using (Nat.find_min _ (lt_of_succ_le hk)) - · apply Finite.not_iff_forall.mp ‹_› + · apply FiniteMultiplicity.not_iff_forall.mp ‹_› theorem pow_dvd_of_le_multiplicity {k : ℕ} (hk : k ≤ multiplicity a b) : a ^ k ∣ b := pow_dvd_of_le_emultiplicity (le_emultiplicity_of_le_multiplicity hk) @@ -220,15 +272,19 @@ theorem not_pow_dvd_of_emultiplicity_lt {m : ℕ} (hm : emultiplicity a b < m) : exact hn2 ((pow_dvd_pow _ hn1).trans nh) · simp at hm -theorem multiplicity.Finite.not_pow_dvd_of_multiplicity_lt (hf : Finite a b) {m : ℕ} +theorem FiniteMultiplicity.not_pow_dvd_of_multiplicity_lt (hf : FiniteMultiplicity a b) {m : ℕ} (hm : multiplicity a b < m) : ¬a ^ m ∣ b := by apply not_pow_dvd_of_emultiplicity_lt rw [hf.emultiplicity_eq_multiplicity] norm_cast +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.not_pow_dvd_of_multiplicity_lt := + FiniteMultiplicity.not_pow_dvd_of_multiplicity_lt + theorem multiplicity_pos_of_dvd (hdiv : a ∣ b) : 0 < multiplicity a b := by refine zero_lt_iff.2 fun h => ?_ - simpa [hdiv] using Finite.not_pow_dvd_of_multiplicity_lt + simpa [hdiv] using FiniteMultiplicity.not_pow_dvd_of_multiplicity_lt (by by_contra! nh; simp [nh] at h) (lt_one_iff.mpr h) theorem emultiplicity_pos_of_dvd (hdiv : a ∣ b) : 0 < emultiplicity a b := @@ -236,7 +292,7 @@ theorem emultiplicity_pos_of_dvd (hdiv : a ∣ b) : 0 < emultiplicity a b := theorem emultiplicity_eq_of_dvd_of_not_dvd {k : ℕ} (hk : a ^ k ∣ b) (hsucc : ¬a ^ (k + 1) ∣ b) : emultiplicity a b = k := by classical - have : Finite a b := ⟨k, hsucc⟩ + have : FiniteMultiplicity a b := ⟨k, hsucc⟩ simp only [emultiplicity, this, ↓reduceDIte, Nat.cast_inj, find_eq_iff, hsucc, not_false_eq_true, Decidable.not_not, true_and] exact fun n hn ↦ (pow_dvd_pow _ hn).trans hk @@ -249,24 +305,36 @@ theorem le_emultiplicity_of_pow_dvd {k : ℕ} (hk : a ^ k ∣ b) : k ≤ emultiplicity a b := le_of_not_gt fun hk' => not_pow_dvd_of_emultiplicity_lt hk' hk -theorem multiplicity.Finite.le_multiplicity_of_pow_dvd (hf : Finite a b) {k : ℕ} (hk : a ^ k ∣ b) : - k ≤ multiplicity a b := +theorem FiniteMultiplicity.le_multiplicity_of_pow_dvd (hf : FiniteMultiplicity a b) + {k : ℕ} (hk : a ^ k ∣ b) : k ≤ multiplicity a b := hf.le_multiplicity_of_le_emultiplicity (le_emultiplicity_of_pow_dvd hk) +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.le_multiplicity_of_pow_dvd := + FiniteMultiplicity.le_multiplicity_of_pow_dvd + theorem pow_dvd_iff_le_emultiplicity {k : ℕ} : a ^ k ∣ b ↔ k ≤ emultiplicity a b := ⟨le_emultiplicity_of_pow_dvd, pow_dvd_of_le_emultiplicity⟩ -theorem multiplicity.Finite.pow_dvd_iff_le_multiplicity (hf : Finite a b) {k : ℕ} : +theorem FiniteMultiplicity.pow_dvd_iff_le_multiplicity (hf : FiniteMultiplicity a b) {k : ℕ} : a ^ k ∣ b ↔ k ≤ multiplicity a b := by exact_mod_cast hf.emultiplicity_eq_multiplicity ▸ pow_dvd_iff_le_emultiplicity +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.pow_dvd_iff_le_multiplicity := + FiniteMultiplicity.pow_dvd_iff_le_multiplicity + theorem emultiplicity_lt_iff_not_dvd {k : ℕ} : emultiplicity a b < k ↔ ¬a ^ k ∣ b := by rw [pow_dvd_iff_le_emultiplicity, not_le] -theorem multiplicity.Finite.multiplicity_lt_iff_not_dvd {k : ℕ} (hf : Finite a b) : +theorem FiniteMultiplicity.multiplicity_lt_iff_not_dvd {k : ℕ} (hf : FiniteMultiplicity a b) : multiplicity a b < k ↔ ¬a ^ k ∣ b := by rw [hf.pow_dvd_iff_le_multiplicity, not_le] +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.multiplicity_lt_iff_not_dvd := + FiniteMultiplicity.multiplicity_lt_iff_not_dvd + theorem emultiplicity_eq_coe {n : ℕ} : emultiplicity a b = n ↔ a ^ n ∣ b ∧ ¬a ^ (n + 1) ∣ b := by constructor @@ -281,33 +349,48 @@ theorem emultiplicity_eq_coe {n : ℕ} : · rw [and_imp] apply emultiplicity_eq_of_dvd_of_not_dvd -theorem multiplicity.Finite.multiplicity_eq_iff (hf : Finite a b) {n : ℕ} : +theorem FiniteMultiplicity.multiplicity_eq_iff (hf : FiniteMultiplicity a b) {n : ℕ} : multiplicity a b = n ↔ a ^ n ∣ b ∧ ¬a ^ (n + 1) ∣ b := by simp [← emultiplicity_eq_coe, hf.emultiplicity_eq_multiplicity] +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.multiplicity_eq_iff := FiniteMultiplicity.multiplicity_eq_iff + @[simp] -theorem multiplicity.Finite.not_of_isUnit_left (b : α) (ha : IsUnit a) : ¬Finite a b := +theorem FiniteMultiplicity.not_of_isUnit_left (b : α) (ha : IsUnit a) : ¬FiniteMultiplicity a b := (·.not_unit ha) -theorem multiplicity.Finite.not_of_one_left (b : α) : ¬ Finite 1 b := by simp +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.not_of_isUnit_left := FiniteMultiplicity.not_of_isUnit_left + +theorem FiniteMultiplicity.not_of_one_left (b : α) : ¬ FiniteMultiplicity 1 b := by simp + +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.not_of_one_left := FiniteMultiplicity.not_of_one_left @[simp] theorem emultiplicity_one_left (b : α) : emultiplicity 1 b = ⊤ := - emultiplicity_eq_top.2 (Finite.not_of_one_left _) + emultiplicity_eq_top.2 (FiniteMultiplicity.not_of_one_left _) @[simp] -theorem multiplicity.Finite.one_right (ha : Finite a 1) : multiplicity a 1 = 0 := by +theorem FiniteMultiplicity.one_right (ha : FiniteMultiplicity a 1) : multiplicity a 1 = 0 := by simp [ha.multiplicity_eq_iff, ha.not_dvd_of_one_right] -theorem multiplicity.Finite.not_of_unit_left (a : α) (u : αˣ) : ¬ Finite (u : α) a := - Finite.not_of_isUnit_left a u.isUnit +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.one_right := FiniteMultiplicity.one_right + +theorem FiniteMultiplicity.not_of_unit_left (a : α) (u : αˣ) : ¬ FiniteMultiplicity (u : α) a := + FiniteMultiplicity.not_of_isUnit_left a u.isUnit + +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.not_of_unit_left := FiniteMultiplicity.not_of_unit_left theorem emultiplicity_eq_zero : emultiplicity a b = 0 ↔ ¬a ∣ b := by - by_cases hf : Finite a b + by_cases hf : FiniteMultiplicity a b · rw [← ENat.coe_zero, emultiplicity_eq_coe] simp - · simpa [emultiplicity_eq_top.2 hf] using Finite.not_iff_forall.1 hf 1 + · simpa [emultiplicity_eq_top.2 hf] using FiniteMultiplicity.not_iff_forall.1 hf 1 theorem multiplicity_eq_zero : multiplicity a b = 0 ↔ ¬a ∣ b := @@ -321,7 +404,7 @@ theorem multiplicity_ne_zero : multiplicity a b ≠ 0 ↔ a ∣ b := by simp [multiplicity_eq_zero] -theorem multiplicity.Finite.exists_eq_pow_mul_and_not_dvd (hfin : Finite a b) : +theorem FiniteMultiplicity.exists_eq_pow_mul_and_not_dvd (hfin : FiniteMultiplicity a b) : ∃ c : α, b = a ^ multiplicity a b * c ∧ ¬a ∣ c := by obtain ⟨c, hc⟩ := pow_multiplicity_dvd a b refine ⟨c, hc, ?_⟩ @@ -330,6 +413,10 @@ theorem multiplicity.Finite.exists_eq_pow_mul_and_not_dvd (hfin : Finite a b) : have h₁ : a ^ (multiplicity a b + 1) ∣ b := ⟨k, hc⟩ exact (hfin.multiplicity_eq_iff.1 (by simp)).2 h₁ +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.exists_eq_pow_mul_and_not_dvd := + FiniteMultiplicity.exists_eq_pow_mul_and_not_dvd + theorem emultiplicity_le_emultiplicity_iff {c d : β} : emultiplicity a b ≤ emultiplicity c d ↔ ∀ n : ℕ, a ^ n ∣ b → c ^ n ∣ d := by classical constructor @@ -349,12 +436,17 @@ theorem emultiplicity_le_emultiplicity_iff {c d : β} : simp_all only [not_exists, Decidable.not_not, not_true_eq_false, top_le_iff, dite_eq_right_iff, ENat.coe_ne_top, imp_false, not_false_eq_true, implies_true] -theorem multiplicity.Finite.multiplicity_le_multiplicity_iff {c d : β} (hab : Finite a b) - (hcd : Finite c d) : multiplicity a b ≤ multiplicity c d ↔ ∀ n : ℕ, a ^ n ∣ b → c ^ n ∣ d := by +theorem FiniteMultiplicity.multiplicity_le_multiplicity_iff {c d : β} (hab : FiniteMultiplicity a b) + (hcd : FiniteMultiplicity c d) : + multiplicity a b ≤ multiplicity c d ↔ ∀ n : ℕ, a ^ n ∣ b → c ^ n ∣ d := by rw [← WithTop.coe_le_coe, ENat.some_eq_coe, ← hab.emultiplicity_eq_multiplicity, ← hcd.emultiplicity_eq_multiplicity] apply emultiplicity_le_emultiplicity_iff +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.multiplicity_le_multiplicity_iff := + FiniteMultiplicity.multiplicity_le_multiplicity_iff + theorem emultiplicity_eq_emultiplicity_iff {c d : β} : emultiplicity a b = emultiplicity c d ↔ ∀ n : ℕ, a ^ n ∣ b ↔ c ^ n ∣ d := ⟨fun h n => @@ -400,8 +492,8 @@ theorem dvd_iff_multiplicity_pos {a b : α} : 0 < multiplicity a b ↔ a ∣ b : theorem dvd_iff_emultiplicity_pos {a b : α} : 0 < emultiplicity a b ↔ a ∣ b := emultiplicity_pos_iff.trans dvd_iff_multiplicity_pos -theorem Nat.multiplicity_finite_iff {a b : ℕ} : Finite a b ↔ a ≠ 1 ∧ 0 < b := by - rw [← not_iff_not, Finite.not_iff_forall, not_and_or, Ne, Classical.not_not, not_lt, +theorem Nat.finiteMultiplicity_iff {a b : ℕ} : FiniteMultiplicity a b ↔ a ≠ 1 ∧ 0 < b := by + rw [← not_iff_not, FiniteMultiplicity.not_iff_forall, not_and_or, not_ne_iff, not_lt, Nat.le_zero] exact ⟨fun h => @@ -417,6 +509,9 @@ theorem Nat.multiplicity_finite_iff {a b : ℕ} : Finite a b ↔ a ≠ 1 ∧ 0 < not_lt_of_ge (le_of_dvd (Nat.pos_of_ne_zero hb) (h b)) (b.lt_pow_self ha_gt_one), fun h => by cases h <;> simp [*]⟩ +@[deprecated (since := "2024-11-30")] +alias Nat.multiplicity_finite_iff := Nat.finiteMultiplicity_iff + alias ⟨_, Dvd.multiplicity_pos⟩ := dvd_iff_multiplicity_pos end Monoid @@ -425,8 +520,11 @@ section CommMonoid variable [CommMonoid α] -theorem multiplicity.Finite.mul_right {a b c : α} (hf : Finite a (b * c)) : Finite a c := - (mul_comm b c ▸ hf).mul_left +theorem FiniteMultiplicity.mul_right {a b c : α} (hf : FiniteMultiplicity a (b * c)) : + FiniteMultiplicity a c := (mul_comm b c ▸ hf).mul_left + +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.mul_right := FiniteMultiplicity.mul_right theorem emultiplicity_of_isUnit_right {a b : α} (ha : ¬IsUnit a) (hb : IsUnit b) : emultiplicity a b = 0 := @@ -471,10 +569,13 @@ section MonoidWithZero variable [MonoidWithZero α] -theorem multiplicity.Finite.ne_zero {a b : α} (h : Finite a b) : b ≠ 0 := +theorem FiniteMultiplicity.ne_zero {a b : α} (h : FiniteMultiplicity a b) : b ≠ 0 := let ⟨n, hn⟩ := h fun hb => by simp [hb] at hn +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.ne_zero := FiniteMultiplicity.ne_zero + @[simp] theorem emultiplicity_zero (a : α) : emultiplicity a 0 = ⊤ := emultiplicity_eq_top.2 (fun v ↦ v.ne_zero rfl) @@ -493,12 +594,15 @@ section Semiring variable [Semiring α] -theorem multiplicity.Finite.or_of_add {p a b : α} (hf : Finite p (a + b)) : - Finite p a ∨ Finite p b := by +theorem FiniteMultiplicity.or_of_add {p a b : α} (hf : FiniteMultiplicity p (a + b)) : + FiniteMultiplicity p a ∨ FiniteMultiplicity p b := by by_contra! nh obtain ⟨c, hc⟩ := hf simp_all [dvd_add] +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.or_of_add := FiniteMultiplicity.or_of_add + theorem min_le_emultiplicity_add {p a b : α} : min (emultiplicity p a) (emultiplicity p b) ≤ emultiplicity p (a + b) := by cases hm : min (emultiplicity p a) (emultiplicity p b) @@ -516,12 +620,19 @@ section Ring variable [Ring α] @[simp] -theorem multiplicity.Finite.neg_iff {a b : α} : Finite a (-b) ↔ Finite a b := by - unfold Finite +theorem FiniteMultiplicity.neg_iff {a b : α} : + FiniteMultiplicity a (-b) ↔ FiniteMultiplicity a b := by + unfold FiniteMultiplicity congr! 3 simp only [dvd_neg] -alias ⟨_, multiplicity.Finite.neg⟩ := Finite.neg_iff +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.neg_iff := FiniteMultiplicity.neg_iff + +alias ⟨_, FiniteMultiplicity.neg⟩ := FiniteMultiplicity.neg_iff + +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.neg := FiniteMultiplicity.neg @[simp] theorem emultiplicity_neg (a b : α) : emultiplicity a (-b) = emultiplicity a b := by @@ -544,7 +655,7 @@ theorem Int.multiplicity_natAbs (a : ℕ) (b : ℤ) : theorem emultiplicity_add_of_gt {p a b : α} (h : emultiplicity p b < emultiplicity p a) : emultiplicity p (a + b) = emultiplicity p b := by - have : Finite p b := finite_iff_emultiplicity_ne_top.2 (by simp [·] at h) + have : FiniteMultiplicity p b := finiteMultiplicity_iff_emultiplicity_ne_top.2 (by simp [·] at h) rw [this.emultiplicity_eq_multiplicity] at * apply emultiplicity_eq_of_dvd_of_not_dvd · apply dvd_add @@ -557,18 +668,21 @@ theorem emultiplicity_add_of_gt {p a b : α} (h : emultiplicity p b < emultiplic apply pow_dvd_of_le_emultiplicity exact Order.add_one_le_of_lt h -theorem multiplicity.Finite.multiplicity_add_of_gt {p a b : α} (hf : Finite p b) +theorem FiniteMultiplicity.multiplicity_add_of_gt {p a b : α} (hf : FiniteMultiplicity p b) (h : multiplicity p b < multiplicity p a) : multiplicity p (a + b) = multiplicity p b := multiplicity_eq_of_emultiplicity_eq <| emultiplicity_add_of_gt (hf.emultiplicity_eq_multiplicity ▸ (WithTop.coe_strictMono h).trans_le multiplicity_le_emultiplicity) +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.multiplicity_add_of_gt := FiniteMultiplicity.multiplicity_add_of_gt + theorem emultiplicity_sub_of_gt {p a b : α} (h : emultiplicity p b < emultiplicity p a) : emultiplicity p (a - b) = emultiplicity p b := by rw [sub_eq_add_neg, emultiplicity_add_of_gt] <;> rw [emultiplicity_neg]; assumption theorem multiplicity_sub_of_gt {p a b : α} (h : multiplicity p b < multiplicity p a) - (hfin : Finite p b) : multiplicity p (a - b) = multiplicity p b := by + (hfin : FiniteMultiplicity p b) : multiplicity p (a - b) = multiplicity p b := by rw [sub_eq_add_neg, hfin.neg.multiplicity_add_of_gt] <;> rw [multiplicity_neg]; assumption theorem emultiplicity_add_eq_min {p a b : α} @@ -581,8 +695,8 @@ theorem emultiplicity_add_eq_min {p a b : α} · rw [emultiplicity_add_of_gt hab, min_eq_right] exact le_of_lt hab -theorem multiplicity_add_eq_min {p a b : α} (ha : Finite p a) (hb : Finite p b) - (h : multiplicity p a ≠ multiplicity p b) : +theorem multiplicity_add_eq_min {p a b : α} (ha : FiniteMultiplicity p a) + (hb : FiniteMultiplicity p b) (h : multiplicity p a ≠ multiplicity p b) : multiplicity p (a + b) = min (multiplicity p a) (multiplicity p b) := by rcases lt_trichotomy (multiplicity p a) (multiplicity p b) with (hab | _ | hab) · rw [add_comm, ha.multiplicity_add_of_gt hab, min_eq_left] @@ -599,7 +713,7 @@ variable [CancelCommMonoidWithZero α] /- Porting note: Pulled a b intro parameters since Lean parses that more easily -/ -theorem multiplicity.finite_mul_aux {p : α} (hp : Prime p) {a b : α} : +theorem finiteMultiplicity_mul_aux {p : α} (hp : Prime p) {a b : α} : ∀ {n m : ℕ}, ¬p ^ (n + 1) ∣ a → ¬p ^ (m + 1) ∣ b → ¬p ^ (n + m + 1) ∣ a * b | n, m => fun ha hb ⟨s, hs⟩ => have : p ∣ a * b := ⟨p ^ (n + m) * s, by simp [hs, pow_add, mul_comm, mul_assoc, mul_left_comm]⟩ @@ -612,7 +726,7 @@ theorem multiplicity.finite_mul_aux {p : α} (hp : Prime p) {a b : α} : rw [tsub_add_cancel_of_le (succ_le_of_lt hn0)] at hy simp [hy, pow_add, mul_comm, mul_assoc, mul_left_comm]⟩) have : 1 ≤ n + m := le_trans hn0 (Nat.le_add_right n m) - finite_mul_aux hp hpx hb + finiteMultiplicity_mul_aux hp hpx hb ⟨s, mul_right_cancel₀ hp.1 (by rw [tsub_add_eq_add_tsub (succ_le_of_lt hn0), tsub_add_cancel_of_le this] simp_all [mul_comm, mul_assoc, mul_left_comm, pow_add])⟩) @@ -626,29 +740,40 @@ theorem multiplicity.finite_mul_aux {p : α} (hp : Prime p) {a b : α} : mul_right_cancel₀ hp.1 <| by rw [tsub_add_cancel_of_le (succ_le_of_lt hm0)] at hy simp [hy, pow_add, mul_comm, mul_assoc, mul_left_comm]⟩) - finite_mul_aux hp ha hpx + finiteMultiplicity_mul_aux hp ha hpx ⟨s, mul_right_cancel₀ hp.1 (by rw [add_assoc, tsub_add_cancel_of_le (succ_le_of_lt hm0)] simp_all [mul_comm, mul_assoc, mul_left_comm, pow_add])⟩ -theorem Prime.multiplicity_finite_mul {p a b : α} (hp : Prime p) : - Finite p a → Finite p b → Finite p (a * b) := - fun ⟨n, hn⟩ ⟨m, hm⟩ => ⟨n + m, finite_mul_aux hp hn hm⟩ +@[deprecated (since := "2024-11-30")] +alias multiplicity.finite_mul_aux := finiteMultiplicity_mul_aux -theorem multiplicity.Finite.mul_iff {p a b : α} (hp : Prime p) : - Finite p (a * b) ↔ Finite p a ∧ Finite p b := +theorem Prime.finiteMultiplicity_mul {p a b : α} (hp : Prime p) : + FiniteMultiplicity p a → FiniteMultiplicity p b → FiniteMultiplicity p (a * b) := + fun ⟨n, hn⟩ ⟨m, hm⟩ => ⟨n + m, finiteMultiplicity_mul_aux hp hn hm⟩ + +@[deprecated (since := "2024-11-30")] +alias Prime.multiplicity_finite_mul := Prime.finiteMultiplicity_mul + +theorem FiniteMultiplicity.mul_iff {p a b : α} (hp : Prime p) : + FiniteMultiplicity p (a * b) ↔ FiniteMultiplicity p a ∧ FiniteMultiplicity p b := ⟨fun h => ⟨h.mul_left, h.mul_right⟩, fun h => - hp.multiplicity_finite_mul h.1 h.2⟩ + hp.finiteMultiplicity_mul h.1 h.2⟩ + +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.mul_iff := FiniteMultiplicity.mul_iff -theorem multiplicity.Finite.pow {p a : α} (hp : Prime p) - (hfin : Finite p a) {k : ℕ} : Finite p (a ^ k) := +theorem FiniteMultiplicity.pow {p a : α} (hp : Prime p) + (hfin : FiniteMultiplicity p a) {k : ℕ} : FiniteMultiplicity p (a ^ k) := match k, hfin with | 0, _ => ⟨0, by simp [mt isUnit_iff_dvd_one.2 hp.2.1]⟩ - | k + 1, ha => by rw [_root_.pow_succ']; exact hp.multiplicity_finite_mul ha (ha.pow hp) + | k + 1, ha => by rw [_root_.pow_succ']; exact hp.finiteMultiplicity_mul ha (ha.pow hp) + +@[deprecated (since := "2024-11-30")] alias multiplicity.Finite.pow := FiniteMultiplicity.pow @[simp] theorem multiplicity_self {a : α} : multiplicity a a = 1 := by - by_cases ha : Finite a a + by_cases ha : FiniteMultiplicity a a · rw [ha.multiplicity_eq_iff] simp only [pow_one, dvd_refl, reduceAdd, true_and] rintro ⟨v, hv⟩ @@ -661,11 +786,14 @@ theorem multiplicity_self {a : α} : multiplicity a a = 1 := by · simp [ha] @[simp] -theorem multiplicity.Finite.emultiplicity_self {a : α} (hfin : Finite a a) : +theorem FiniteMultiplicity.emultiplicity_self {a : α} (hfin : FiniteMultiplicity a a) : emultiplicity a a = 1 := by simp [hfin.emultiplicity_eq_multiplicity] -theorem multiplicity_mul {p a b : α} (hp : Prime p) (hfin : Finite p (a * b)) : +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.emultiplicity_self := FiniteMultiplicity.emultiplicity_self + +theorem multiplicity_mul {p a b : α} (hp : Prime p) (hfin : FiniteMultiplicity p (a * b)) : multiplicity p (a * b) = multiplicity p a + multiplicity p b := by have hdiva : p ^ multiplicity p a ∣ a := pow_multiplicity_dvd .. have hdivb : p ^ multiplicity p b ∣ b := pow_multiplicity_dvd .. @@ -681,14 +809,14 @@ theorem multiplicity_mul {p a b : α} (hp : Prime p) (hfin : Finite p (a * b)) : theorem emultiplicity_mul {p a b : α} (hp : Prime p) : emultiplicity p (a * b) = emultiplicity p a + emultiplicity p b := by - by_cases hfin : Finite p (a * b) + by_cases hfin : FiniteMultiplicity p (a * b) · rw [hfin.emultiplicity_eq_multiplicity, hfin.mul_left.emultiplicity_eq_multiplicity, hfin.mul_right.emultiplicity_eq_multiplicity] norm_cast exact multiplicity_mul hp hfin · rw [emultiplicity_eq_top.2 hfin, eq_comm, WithTop.add_eq_top, emultiplicity_eq_top, emultiplicity_eq_top] - simpa only [Finite.mul_iff hp, not_and_or] using hfin + simpa only [FiniteMultiplicity.mul_iff hp, not_and_or] using hfin theorem Finset.emultiplicity_prod {β : Type*} {p : α} (hp : Prime p) (s : Finset β) (f : β → α) : emultiplicity p (∏ x ∈ s, f x) = ∑ x ∈ s, emultiplicity p (f x) := by classical @@ -703,11 +831,14 @@ theorem emultiplicity_pow {p a : α} (hp : Prime p) {k : ℕ} : · simp [emultiplicity_of_one_right hp.not_unit] · simp [pow_succ, emultiplicity_mul hp, hk, add_mul] -protected theorem multiplicity.Finite.multiplicity_pow {p a : α} (hp : Prime p) - (ha : Finite p a) {k : ℕ} : multiplicity p (a ^ k) = k * multiplicity p a := by +protected theorem FiniteMultiplicity.multiplicity_pow {p a : α} (hp : Prime p) + (ha : FiniteMultiplicity p a) {k : ℕ} : multiplicity p (a ^ k) = k * multiplicity p a := by exact_mod_cast (ha.pow hp).emultiplicity_eq_multiplicity ▸ ha.emultiplicity_eq_multiplicity ▸ emultiplicity_pow hp +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.multiplicity_pow := FiniteMultiplicity.multiplicity_pow + theorem emultiplicity_pow_self {p : α} (h0 : p ≠ 0) (hu : ¬IsUnit p) (n : ℕ) : emultiplicity p (p ^ n) = n := by apply emultiplicity_eq_of_dvd_of_not_dvd @@ -731,8 +862,6 @@ end CancelCommMonoidWithZero section Nat -open multiplicity - theorem multiplicity_eq_zero_of_coprime {p a b : ℕ} (hp : p ≠ 1) (hle : multiplicity p a ≤ multiplicity p b) (hab : Nat.Coprime a b) : multiplicity p a = 0 := by apply Nat.eq_zero_of_not_pos @@ -745,16 +874,29 @@ theorem multiplicity_eq_zero_of_coprime {p a b : ℕ} (hp : p ≠ 1) end Nat -theorem Int.multiplicity_finite_iff_natAbs_finite {a b : ℤ} : - Finite a b ↔ Finite a.natAbs b.natAbs := by - simp only [Finite.def, ← Int.natAbs_dvd_natAbs, Int.natAbs_pow] +theorem Int.finiteMultiplicity_iff_finiteMultiplicity_natAbs {a b : ℤ} : + FiniteMultiplicity a b ↔ FiniteMultiplicity a.natAbs b.natAbs := by + simp only [FiniteMultiplicity.def, ← Int.natAbs_dvd_natAbs, Int.natAbs_pow] + +@[deprecated (since := "2024-11-30")] +alias Int.multiplicity_finite_iff_natAbs_finite := + Int.finiteMultiplicity_iff_finiteMultiplicity_natAbs -theorem Int.multiplicity_finite_iff {a b : ℤ} : Finite a b ↔ a.natAbs ≠ 1 ∧ b ≠ 0 := by - rw [multiplicity_finite_iff_natAbs_finite, Nat.multiplicity_finite_iff, +theorem Int.finiteMultiplicity_iff {a b : ℤ} : FiniteMultiplicity a b ↔ a.natAbs ≠ 1 ∧ b ≠ 0 := by + rw [finiteMultiplicity_iff_finiteMultiplicity_natAbs, Nat.finiteMultiplicity_iff, pos_iff_ne_zero, Int.natAbs_ne_zero] -instance Nat.decidableMultiplicityFinite : DecidableRel fun a b : ℕ => Finite a b := fun _ _ => - decidable_of_iff' _ Nat.multiplicity_finite_iff +@[deprecated (since := "2024-11-30")] +alias Int.multiplicity_finite_iff := Int.finiteMultiplicity_iff + +instance Nat.decidableFiniteMultiplicity : DecidableRel fun a b : ℕ => FiniteMultiplicity a b := + fun _ _ ↦ decidable_of_iff' _ Nat.finiteMultiplicity_iff + +@[deprecated (since := "2024-11-30")] +alias Nat.decidableMultiplicityFinite := Nat.decidableFiniteMultiplicity + +instance Int.decidableMultiplicityFinite : DecidableRel fun a b : ℤ => FiniteMultiplicity a b := + fun _ _ ↦ decidable_of_iff' _ Int.finiteMultiplicity_iff -instance Int.decidableMultiplicityFinite : DecidableRel fun a b : ℤ => Finite a b := fun _ _ => - decidable_of_iff' _ Int.multiplicity_finite_iff +@[deprecated (since := "2024-11-30")] +alias Int.decidableFiniteMultiplicity := Int.decidableMultiplicityFinite diff --git a/Mathlib/RingTheory/PowerSeries/Order.lean b/Mathlib/RingTheory/PowerSeries/Order.lean index 70348c70fd666..381b894105712 100644 --- a/Mathlib/RingTheory/PowerSeries/Order.lean +++ b/Mathlib/RingTheory/PowerSeries/Order.lean @@ -38,8 +38,6 @@ variable {R : Type*} section OrderBasic -open multiplicity - variable [Semiring R] {φ : R⟦X⟧} theorem exists_coeff_ne_zero_iff_ne_zero : (∃ n : ℕ, coeff R n φ ≠ 0) ↔ φ ≠ 0 := by diff --git a/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean b/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean index df43e22b8bed7..2da9825d08c25 100644 --- a/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean +++ b/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean @@ -150,7 +150,7 @@ theorem minpoly_eq_pow {p : ℕ} [hprime : Fact p.Prime] (hdiv : ¬p ∣ n) : (separable_X_pow_sub_C 1 (fun h => hdiv <| (ZMod.natCast_zmod_eq_zero_iff_dvd n p).1 h) one_ne_zero).squarefree cases' - (multiplicity.squarefree_iff_emultiplicity_le_one (X ^ n - 1)).1 hfree + (squarefree_iff_emultiplicity_le_one (X ^ n - 1)).1 hfree (map (Int.castRingHom (ZMod p)) P) with hle hunit · rw [Nat.cast_one] at habs; exact hle.not_lt habs diff --git a/Mathlib/RingTheory/RootsOfUnity/PrimitiveRoots.lean b/Mathlib/RingTheory/RootsOfUnity/PrimitiveRoots.lean index 8399113e64d64..f1f8616f7cb18 100644 --- a/Mathlib/RingTheory/RootsOfUnity/PrimitiveRoots.lean +++ b/Mathlib/RingTheory/RootsOfUnity/PrimitiveRoots.lean @@ -363,7 +363,7 @@ theorem primitiveRoots_one : primitiveRoots 1 R = {(1 : R)} := by theorem neZero' {n : ℕ} [NeZero n] (hζ : IsPrimitiveRoot ζ n) : NeZero ((n : ℕ) : R) := by let p := ringChar R - have hfin := Nat.multiplicity_finite_iff.2 ⟨CharP.char_ne_one R p, NeZero.pos n⟩ + have hfin := Nat.finiteMultiplicity_iff.2 ⟨CharP.char_ne_one R p, NeZero.pos n⟩ obtain ⟨m, hm⟩ := hfin.exists_eq_pow_mul_and_not_dvd by_cases hp : p ∣ n · obtain ⟨k, hk⟩ := Nat.exists_eq_succ_of_ne_zero (multiplicity_pos_of_dvd hp).ne' diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean index dbc52cb56fad1..a6504ae1d7bda 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean @@ -33,11 +33,21 @@ theorem WfDvdMonoid.max_power_factor [CommMonoidWithZero α] [WfDvdMonoid α] {a (h : a₀ ≠ 0) (hx : Irreducible x) : ∃ (n : ℕ) (a : α), ¬x ∣ a ∧ a₀ = x ^ n * a := max_power_factor' h hx.not_unit -theorem multiplicity.finite_of_not_isUnit [CancelCommMonoidWithZero α] [WfDvdMonoid α] - {a b : α} (ha : ¬IsUnit a) (hb : b ≠ 0) : multiplicity.Finite a b := by +theorem FiniteMultiplicity.of_not_isUnit [CancelCommMonoidWithZero α] [WfDvdMonoid α] + {a b : α} (ha : ¬IsUnit a) (hb : b ≠ 0) : FiniteMultiplicity a b := by obtain ⟨n, c, ndvd, rfl⟩ := WfDvdMonoid.max_power_factor' hb ha exact ⟨n, by rwa [pow_succ, mul_dvd_mul_iff_left (left_ne_zero_of_mul hb)]⟩ +@[deprecated (since := "2024-11-30")] +alias multiplicity.finite_of_not_isUnit := FiniteMultiplicity.of_not_isUnit + +theorem FiniteMultiplicity.of_prime_left [CancelCommMonoidWithZero α] [WfDvdMonoid α] + {a b : α} (ha : Prime a) (hb : b ≠ 0) : FiniteMultiplicity a b := + .of_not_isUnit ha.not_unit hb + +@[deprecated (since := "2024-11-30")] +alias multiplicity.finite_prime_left := FiniteMultiplicity.of_prime_left + namespace UniqueFactorizationMonoid variable {R : Type*} [CancelCommMonoidWithZero R] [UniqueFactorizationMonoid R] From 1957fef0f098a197410617514519012cccddd207 Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Fri, 6 Dec 2024 15:26:42 +0000 Subject: [PATCH 551/829] feat(CategoryTheory): expand the API for AB axioms (#19200) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR introduces typeclasses `HasExactColimitsOfShape J C` and `HasExactLimitsOfShape J C` which say that a `J`-shaped colimits (resp. limits) in a category `C` are exact in the sense of commuting with finite limits (resp. colimits). We redefine `AB4`, `AB4Star`, etc. in terms of these, and add also countable versions. Further, we generalize the proof of `AB5` implies `AB4` to a statement about `HasExactColimitsOfShape`, and add API for transporting instances of these classes along final/initial functors. --- Mathlib/Algebra/Category/Grp/AB5.lean | 9 +- .../Abelian/GrothendieckAxioms.lean | 420 ++++++++++++++++-- Mathlib/CategoryTheory/Limits/Filtered.lean | 28 +- .../Limits/Shapes/Products.lean | 13 +- 4 files changed, 419 insertions(+), 51 deletions(-) diff --git a/Mathlib/Algebra/Category/Grp/AB5.lean b/Mathlib/Algebra/Category/Grp/AB5.lean index a5230f8bf90ff..81a72475dbf30 100644 --- a/Mathlib/Algebra/Category/Grp/AB5.lean +++ b/Mathlib/Algebra/Category/Grp/AB5.lean @@ -3,14 +3,9 @@ Copyright (c) 2023 David Kurniadi Angdinata. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Kurniadi Angdinata, Moritz Firsching, Nikolas Kuhn, Amelia Livingston -/ - -import Mathlib.Algebra.Homology.ShortComplex.Ab -import Mathlib.Algebra.Homology.ShortComplex.ExactFunctor -import Mathlib.CategoryTheory.Abelian.Exact import Mathlib.Algebra.Category.Grp.FilteredColimits -import Mathlib.CategoryTheory.Abelian.FunctorCategory +import Mathlib.Algebra.Homology.ShortComplex.Ab import Mathlib.CategoryTheory.Abelian.GrothendieckAxioms - /-! # The category of abelian groups satisfies Grothendieck's axiom AB5 @@ -51,4 +46,4 @@ instance : HasFilteredColimits (AddCommGrp.{u}) where HasColimitsOfShape := inferInstance noncomputable instance : AB5 (AddCommGrp.{u}) where - preservesFiniteLimits := fun _ => inferInstance + ofShape _ := { preservesFiniteLimits := inferInstance } diff --git a/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms.lean b/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms.lean index e11ac111c917c..9e81a28ffd8e2 100644 --- a/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms.lean +++ b/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms.lean @@ -3,11 +3,12 @@ Copyright (c) 2023 Adam Topaz. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Isaac Hernando, Coleton Kotch, Adam Topaz -/ - +import Mathlib.Algebra.Homology.ShortComplex.ExactFunctor +import Mathlib.CategoryTheory.Abelian.FunctorCategory import Mathlib.CategoryTheory.Limits.Constructions.Filtered -import Mathlib.CategoryTheory.Limits.Shapes.Biproducts import Mathlib.CategoryTheory.Limits.Preserves.FunctorCategory - +import Mathlib.CategoryTheory.Limits.Shapes.Countable +import Mathlib.Logic.Equiv.List /-! # Grothendieck Axioms @@ -17,9 +18,10 @@ basic facts about them. ## Definitions -- `AB4` -- a category satisfies `AB4` provided that coproducts are exact. -- `AB5` -- a category satisfies `AB5` provided that filtered colimits are exact. -- The duals of the above definitions, called `AB4Star` and `AB5Star`. +- `HasExactColimitsOfShape J` -- colimits of shape `J` are exact. +- The dual of the above definitions, called `HasExactLimitsOfShape`. +- `AB4` -- coproducts are exact (this is formulated in terms of `HasExactColimitsOfShape`). +- `AB5` -- filtered colimits are exact (this is formulated in terms of `HasExactColimitsOfShape`). ## Theorems @@ -50,63 +52,226 @@ namespace CategoryTheory open Limits -universe v v' u u' w +attribute [instance] comp_preservesFiniteLimits comp_preservesFiniteColimits + +universe w w' w₂ w₂' v v' u u' variable (C : Type u) [Category.{v} C] -attribute [local instance] hasCoproducts_of_finite_and_filtered +/-- +A category `C` is said to have exact colimits of shape `J` provided that colimits of shape `J` +exist and are exact (in the sense that they preserve finite limits). +-/ +class HasExactColimitsOfShape (J : Type u') [Category.{v'} J] (C : Type u) [Category.{v} C] + [HasColimitsOfShape J C] where + /-- Exactness of `J`-shaped colimits stated as `colim : (J ⥤ C) ⥤ C` preserving finite limits. -/ + preservesFiniteLimits : PreservesFiniteLimits (colim (J := J) (C := C)) + +/-- +A category `C` is said to have exact limits of shape `J` provided that limits of shape `J` +exist and are exact (in the sense that they preserve finite colimits). +-/ +class HasExactLimitsOfShape (J : Type u') [Category.{v'} J] (C : Type u) [Category.{v} C] + [HasLimitsOfShape J C] where + /-- Exactness of `J`-shaped limits stated as `lim : (J ⥤ C) ⥤ C` preserving finite colimits. -/ + preservesFiniteColimits : PreservesFiniteColimits (lim (J := J) (C := C)) + +attribute [instance] HasExactColimitsOfShape.preservesFiniteLimits + HasExactLimitsOfShape.preservesFiniteColimits + +/-- +Transport a `HasExactColimitsOfShape` along an equivalence of the shape. + +Note: When `C` has finite limits, this lemma holds with the equivalence replaced by a final +functor, see `hasExactColimitsOfShape_of_final` below. +-/ +lemma hasExactColimitsOfShape_of_equiv {J J' : Type*} [Category J] [Category J'] (e : J ≌ J') + [HasColimitsOfShape J C] [HasExactColimitsOfShape J C] : + haveI : HasColimitsOfShape J' C := hasColimitsOfShape_of_equivalence e + HasExactColimitsOfShape J' C := + haveI : HasColimitsOfShape J' C := hasColimitsOfShape_of_equivalence e + ⟨preservesFiniteLimits_of_natIso (Functor.Final.colimIso e.functor)⟩ + +/-- +Transport a `HasExactLimitsOfShape` along an equivalence of the shape. + +Note: When `C` has finite colimits, this lemma holds with the equivalence replaced by a initial +functor, see `hasExactLimitsOfShape_of_initial` below. +-/ +lemma hasExactLimitsOfShape_of_equiv {J J' : Type*} [Category J] [Category J'] (e : J ≌ J') + [HasLimitsOfShape J C] [HasExactLimitsOfShape J C] : + haveI : HasLimitsOfShape J' C := hasLimitsOfShape_of_equivalence e + HasExactLimitsOfShape J' C := + haveI : HasLimitsOfShape J' C := hasLimitsOfShape_of_equivalence e + ⟨preservesFiniteColimits_of_natIso (Functor.Initial.limIso e.functor)⟩ + +/-- +A category `C` which has coproducts is said to have `AB4` of size `w` provided that +coproducts of size `w` are exact. +-/ +@[pp_with_univ] +class AB4OfSize [HasCoproducts.{w} C] where + ofShape (α : Type w) : HasExactColimitsOfShape (Discrete α) C + +attribute [instance] AB4OfSize.ofShape /-- A category `C` which has coproducts is said to have `AB4` provided that coproducts are exact. -/ -class AB4 [HasCoproducts C] where - /-- Exactness of coproducts stated as `colim : (Discrete α ⥤ C) ⥤ C` preserving limits. -/ - preservesFiniteLimits (α : Type v) : - PreservesFiniteLimits (colim (J := Discrete α) (C := C)) +abbrev AB4 [HasCoproducts C] := AB4OfSize.{v} C + +lemma AB4OfSize_shrink [HasCoproducts.{max w w'} C] [AB4OfSize.{max w w'} C] : + haveI : HasCoproducts.{w} C := hasCoproducts_shrink.{w, w'} + AB4OfSize.{w} C := + haveI := hasCoproducts_shrink.{w, w'} (C := C) + ⟨fun J ↦ hasExactColimitsOfShape_of_equiv C + (Discrete.equivalence Equiv.ulift : Discrete (ULift.{w'} J) ≌ _)⟩ + +instance (priority := 100) [HasCoproducts.{w} C] [AB4OfSize.{w} C] : + haveI : HasCoproducts.{0} C := hasCoproducts_shrink + AB4OfSize.{0} C := AB4OfSize_shrink C + +/-- A category `C` which has products is said to have `AB4Star` (in literature `AB4*`) +provided that products are exact. -/ +@[pp_with_univ] +class AB4StarOfSize [HasProducts.{w} C] where + ofShape (α : Type w) : HasExactLimitsOfShape (Discrete α) C -attribute [instance] AB4.preservesFiniteLimits +attribute [instance] AB4StarOfSize.ofShape /-- A category `C` which has products is said to have `AB4Star` (in literature `AB4*`) provided that products are exact. -/ -class AB4Star [HasProducts C] where - /-- Exactness of products stated as `lim : (Discrete α ⥤ C) ⥤ C` preserving colimits. -/ - preservesFiniteColimits (α : Type v) : - PreservesFiniteColimits (lim (J := Discrete α) (C := C)) +abbrev AB4Star [HasProducts C] := AB4StarOfSize.{v} C + +lemma AB4StarOfSize_shrink [HasProducts.{max w w'} C] [AB4StarOfSize.{max w w'} C] : + haveI : HasProducts.{w} C := hasProducts_shrink.{w, w'} + AB4StarOfSize.{w} C := + haveI := hasProducts_shrink.{w, w'} (C := C) + ⟨fun J ↦ hasExactLimitsOfShape_of_equiv C + (Discrete.equivalence Equiv.ulift : Discrete (ULift.{w'} J) ≌ _)⟩ + +instance (priority := 100) [HasProducts.{w} C] [AB4StarOfSize.{w} C] : + haveI : HasProducts.{0} C := hasProducts_shrink + AB4StarOfSize.{0} C := AB4StarOfSize_shrink C + +/-- +A category `C` which has countable coproducts is said to have countable `AB4` provided that +countable coproducts are exact. +-/ +class CountableAB4 [HasCountableCoproducts C] where + ofShape (α : Type) [Countable α] : HasExactColimitsOfShape (Discrete α) C + +instance (priority := 100) [HasCoproducts.{0} C] [AB4OfSize.{0} C] : CountableAB4 C := + ⟨inferInstance⟩ + +/-- +A category `C` which has countable coproducts is said to have countable `AB4Star` provided that +countable products are exact. +-/ +class CountableAB4Star [HasCountableProducts C] where + ofShape (α : Type) [Countable α] : HasExactLimitsOfShape (Discrete α) C + +instance (priority := 100) [HasProducts.{0} C] [AB4StarOfSize.{0} C] : CountableAB4Star C := + ⟨inferInstance⟩ -attribute [instance] AB4Star.preservesFiniteColimits +attribute [instance] CountableAB4.ofShape CountableAB4Star.ofShape + +/-- +A category `C` which has filtered colimits of a given size is said to have `AB5` of that size +provided that these filtered colimits are exact. + +`AB5OfSize.{w, w'} C` means that `C` has exact colimits of shape `J : Type w'` with +`Category.{w} J` such that `J` is filtered. +-/ +@[pp_with_univ] +class AB5OfSize [HasFilteredColimitsOfSize.{w, w'} C] where + ofShape (J : Type w') [Category.{w} J] [IsFiltered J] : HasExactColimitsOfShape J C + +attribute [instance] AB5OfSize.ofShape /-- A category `C` which has filtered colimits is said to have `AB5` provided that filtered colimits are exact. -/ -class AB5 [HasFilteredColimits C] where - /-- Exactness of filtered colimits stated as `colim : (J ⥤ C) ⥤ C` on filtered `J` - preserving limits. -/ - preservesFiniteLimits (J : Type v) [SmallCategory J] [IsFiltered J] : - PreservesFiniteLimits (colim (J := J) (C := C)) - -attribute [instance] AB5.preservesFiniteLimits +abbrev AB5 [HasFilteredColimits C] := AB5OfSize.{v, v} C + +lemma AB5OfSize_of_univLE [HasFilteredColimitsOfSize.{w₂, w₂'} C] [UnivLE.{w, w₂}] + [UnivLE.{w', w₂'}] [AB5OfSize.{w₂, w₂'} C] : + haveI : HasFilteredColimitsOfSize.{w, w'} C := hasFilteredColimitsOfSize_of_univLE.{w} + AB5OfSize.{w, w'} C := by + haveI : HasFilteredColimitsOfSize.{w, w'} C := hasFilteredColimitsOfSize_of_univLE.{w} + constructor + intro J _ _ + haveI := IsFiltered.of_equivalence ((ShrinkHoms.equivalence.{w₂} J).trans <| + Shrink.equivalence.{w₂'} (ShrinkHoms.{w'} J)) + exact hasExactColimitsOfShape_of_equiv _ ((ShrinkHoms.equivalence.{w₂} J).trans <| + Shrink.equivalence.{w₂'} (ShrinkHoms.{w'} J)).symm + +lemma AB5OfSize_shrink [HasFilteredColimitsOfSize.{max w w₂, max w' w₂'} C] + [AB5OfSize.{max w w₂, max w' w₂'} C] : + haveI : HasFilteredColimitsOfSize.{w, w'} C := hasFilteredColimitsOfSize_shrink + AB5OfSize.{w, w'} C := + AB5OfSize_of_univLE C /-- A category `C` which has cofiltered limits is said to have `AB5Star` (in literature `AB5*`) provided that cofiltered limits are exact. -/ -class AB5Star [HasCofilteredLimits C] where - /-- Exactness of cofiltered limits stated as `lim : (J ⥤ C) ⥤ C` on cofiltered `J` - preserving colimits. -/ - preservesFiniteColimits (J : Type v) [SmallCategory J] [IsCofiltered J] : - PreservesFiniteColimits (lim (J := J) (C := C)) +@[pp_with_univ] +class AB5StarOfSize [HasCofilteredLimitsOfSize.{w, w'} C] where + ofShape (J : Type w') [Category.{w} J] [IsCofiltered J] : HasExactLimitsOfShape J C -attribute [instance] AB5Star.preservesFiniteColimits +attribute [instance] AB5StarOfSize.ofShape -noncomputable section +/-- +A category `C` which has cofiltered limits is said to have `AB5Star` (in literature `AB5*`) +provided that cofiltered limits are exact. +-/ +abbrev AB5Star [HasCofilteredLimits C] := AB5StarOfSize.{v, v} C + +lemma AB5StarOfSize_of_univLE [HasCofilteredLimitsOfSize.{w₂, w₂'} C] [UnivLE.{w, w₂}] + [UnivLE.{w', w₂'}] [AB5StarOfSize.{w₂, w₂'} C] : + haveI : HasCofilteredLimitsOfSize.{w, w'} C := hasCofilteredLimitsOfSize_of_univLE.{w} + AB5StarOfSize.{w, w'} C := by + haveI : HasCofilteredLimitsOfSize.{w, w'} C := hasCofilteredLimitsOfSize_of_univLE.{w} + constructor + intro J _ _ + haveI := IsCofiltered.of_equivalence ((ShrinkHoms.equivalence.{w₂} J).trans <| + Shrink.equivalence.{w₂'} (ShrinkHoms.{w'} J)) + exact hasExactLimitsOfShape_of_equiv _ ((ShrinkHoms.equivalence.{w₂} J).trans <| + Shrink.equivalence.{w₂'} (ShrinkHoms.{w'} J)).symm + +lemma AB5StarOfSize_shrink [HasCofilteredLimitsOfSize.{max w w₂, max w' w₂'} C] + [AB5StarOfSize.{max w w₂, max w' w₂'} C] : + haveI : HasCofilteredLimitsOfSize.{w, w'} C := hasCofilteredLimitsOfSize_shrink + AB5StarOfSize.{w, w'} C := + AB5StarOfSize_of_univLE C + +/-- `HasExactColimitsOfShape` can be "pushed forward" along final functors -/ +lemma hasExactColimitsOfShape_of_final [HasFiniteLimits C] {J J' : Type*} [Category J] [Category J'] + (F : J ⥤ J') [F.Final] [HasColimitsOfShape J' C] [HasColimitsOfShape J C] + [HasExactColimitsOfShape J C] : HasExactColimitsOfShape J' C where + preservesFiniteLimits := + letI : PreservesFiniteLimits ((whiskeringLeft J J' C).obj F) := ⟨fun _ ↦ inferInstance⟩ + letI := comp_preservesFiniteLimits ((whiskeringLeft J J' C).obj F) colim + preservesFiniteLimits_of_natIso (Functor.Final.colimIso F) + +/-- `HasExactLimitsOfShape` can be "pushed forward" along initial functors -/ +lemma hasExactLimitsOfShape_of_initial [HasFiniteColimits C] {J J' : Type*} [Category J] + [Category J'] (F : J ⥤ J') [F.Initial] [HasLimitsOfShape J' C] [HasLimitsOfShape J C] + [HasExactLimitsOfShape J C] : HasExactLimitsOfShape J' C where + preservesFiniteColimits := + letI : PreservesFiniteColimits ((whiskeringLeft J J' C).obj F) := ⟨fun _ ↦ inferInstance⟩ + letI := comp_preservesFiniteColimits ((whiskeringLeft J J' C).obj F) lim + preservesFiniteColimits_of_natIso (Functor.Initial.limIso F) + +section AB4OfAB5 + +variable {α : Type w} [HasZeroMorphisms C] [HasFiniteBiproducts C] [HasFiniteLimits C] open CoproductsFromFiniteFiltered -variable {α : Type w} -variable [HasZeroMorphisms C] [HasFiniteBiproducts C] [HasFiniteLimits C] - instance preservesFiniteLimits_liftToFinset : PreservesFiniteLimits (liftToFinset C α) := preservesFiniteLimits_of_evaluation _ fun I => letI : PreservesFiniteLimits (colim (J := Discrete I) (C := C)) := @@ -117,15 +282,192 @@ instance preservesFiniteLimits_liftToFinset : PreservesFiniteLimits (liftToFinse letI : PreservesFiniteLimits ((whiskeringLeft (Discrete I) (Discrete α) C).obj (Discrete.functor (·.val)) ⋙ colim) := comp_preservesFiniteLimits _ _ - preservesFiniteLimits_of_natIso (liftToFinsetEvaluationIso I).symm + preservesFiniteLimits_of_natIso (liftToFinsetEvaluationIso I).symm -/-- A category with finite biproducts and finite limits is AB4 if it is AB5. -/ -lemma AB4.of_AB5 [HasFilteredColimits C] [AB5 C] : AB4 C where - preservesFiniteLimits J := +variable (J : Type*) + +/-- +`HasExactColimitsOfShape (Finset (Discrete J)) C` implies `HasExactColimitsOfShape (Discrete J) C` +-/ +lemma hasExactColimitsOfShape_discrete_of_hasExactColimitsOfShape_finset_discrete + [HasColimitsOfShape (Discrete J) C] [HasColimitsOfShape (Finset (Discrete J)) C] + [HasExactColimitsOfShape (Finset (Discrete J)) C] : HasExactColimitsOfShape (Discrete J) C where + preservesFiniteLimits := letI : PreservesFiniteLimits (liftToFinset C J ⋙ colim) := comp_preservesFiniteLimits _ _ preservesFiniteLimits_of_natIso (liftToFinsetColimIso) +attribute [local instance] hasCoproducts_of_finite_and_filtered in +/-- A category with finite biproducts and finite limits is AB4 if it is AB5. -/ +lemma AB4.of_AB5 [HasFilteredColimitsOfSize.{w, w} C] + [AB5OfSize.{w, w} C] : AB4OfSize.{w} C where + ofShape _ := hasExactColimitsOfShape_discrete_of_hasExactColimitsOfShape_finset_discrete _ _ + +/-- +A category with finite biproducts and finite limits has countable AB4 if sequential colimits are +exact. +-/ +lemma CountableAB4.of_countableAB5 [HasColimitsOfShape ℕ C] [HasExactColimitsOfShape ℕ C] + [HasCountableCoproducts C] : CountableAB4 C where + ofShape J := + have : HasColimitsOfShape (Finset (Discrete J)) C := + Functor.Final.hasColimitsOfShape_of_final + (IsFiltered.sequentialFunctor (Finset (Discrete J))) + have := hasExactColimitsOfShape_of_final C (IsFiltered.sequentialFunctor (Finset (Discrete J))) + hasExactColimitsOfShape_discrete_of_hasExactColimitsOfShape_finset_discrete _ _ + +end AB4OfAB5 + +section AB4StarOfAB5Star + +variable {α : Type w} [HasZeroMorphisms C] [HasFiniteBiproducts C] [HasFiniteColimits C] + +open ProductsFromFiniteCofiltered + +instance preservesFiniteColimits_liftToFinset : PreservesFiniteColimits (liftToFinset C α) := + preservesFiniteColimits_of_evaluation _ fun ⟨I⟩ => + letI : PreservesFiniteColimits (lim (J := Discrete I) (C := C)) := + preservesFiniteColimits_of_natIso HasBiproductsOfShape.colimIsoLim + letI : PreservesFiniteColimits ((whiskeringLeft (Discrete I) (Discrete α) C).obj + (Discrete.functor fun x ↦ ↑x)) := ⟨fun _ _ _ => inferInstance⟩ + letI : PreservesFiniteColimits ((whiskeringLeft (Discrete I) (Discrete α) C).obj + (Discrete.functor (·.val)) ⋙ lim) := + comp_preservesFiniteColimits _ _ + preservesFiniteColimits_of_natIso (liftToFinsetEvaluationIso _ _ I).symm + +variable (J : Type*) + +/-- +`HasExactLimitsOfShape (Finset (Discrete J))ᵒᵖ C` implies `HasExactLimitsOfShape (Discrete J) C` +-/ +lemma hasExactLimitsOfShape_discrete_of_hasExactLimitsOfShape_finset_discrete_op + [HasLimitsOfShape (Discrete J) C] [HasLimitsOfShape (Finset (Discrete J))ᵒᵖ C] + [HasExactLimitsOfShape (Finset (Discrete J))ᵒᵖ C] : + HasExactLimitsOfShape (Discrete J) C where + preservesFiniteColimits := + letI : PreservesFiniteColimits (ProductsFromFiniteCofiltered.liftToFinset C J ⋙ lim) := + comp_preservesFiniteColimits _ _ + preservesFiniteColimits_of_natIso (ProductsFromFiniteCofiltered.liftToFinsetLimIso _ _) + +attribute [local instance] hasProducts_of_finite_and_cofiltered in +/-- A category with finite biproducts and finite limits is AB4 if it is AB5. -/ +lemma AB4Star.of_AB5Star [HasCofilteredLimitsOfSize.{w, w} C] [AB5StarOfSize.{w, w} C] : + AB4StarOfSize.{w} C where + ofShape _ := hasExactLimitsOfShape_discrete_of_hasExactLimitsOfShape_finset_discrete_op _ _ + +/-- +A category with finite biproducts and finite limits has countable AB4* if sequential limits are +exact. +-/ +lemma CountableAB4Star.of_countableAB5Star [HasLimitsOfShape ℕᵒᵖ C] [HasExactLimitsOfShape ℕᵒᵖ C] + [HasCountableProducts C] : CountableAB4Star C where + ofShape J := + have : HasLimitsOfShape (Finset (Discrete J))ᵒᵖ C := + Functor.Initial.hasLimitsOfShape_of_initial + (IsFiltered.sequentialFunctor (Finset (Discrete J))).op + have := hasExactLimitsOfShape_of_initial C + (IsFiltered.sequentialFunctor (Finset (Discrete J))).op + hasExactLimitsOfShape_discrete_of_hasExactLimitsOfShape_finset_discrete_op _ _ + +end AB4StarOfAB5Star + +/-- +Checking exactness of colimits of shape `Discrete ℕ` and `Discrete J` for finite `J` is enough for +countable AB4. +-/ +lemma CountableAB4.of_hasExactColimitsOfShape_nat_and_finite [HasCountableCoproducts C] + [HasFiniteLimits C] [∀ (J : Type) [Finite J], HasExactColimitsOfShape (Discrete J) C] + [HasExactColimitsOfShape (Discrete ℕ) C] : + CountableAB4 C where + ofShape J := by + by_cases h : Finite J + · infer_instance + · have : Infinite J := ⟨h⟩ + let _ := Encodable.ofCountable J + let _ := Denumerable.ofEncodableOfInfinite J + exact hasExactColimitsOfShape_of_final C (Discrete.equivalence (Denumerable.eqv J)).inverse + +/-- +Checking exactness of limits of shape `Discrete ℕ` and `Discrete J` for finite `J` is enough for +countable AB4*. +-/ +lemma CountableAB4Star.of_hasExactLimitsOfShape_nat_and_finite [HasCountableProducts C] + [HasFiniteColimits C] [∀ (J : Type) [Finite J], HasExactLimitsOfShape (Discrete J) C] + [HasExactLimitsOfShape (Discrete ℕ) C] : + CountableAB4Star C where + ofShape J := by + by_cases h : Finite J + · infer_instance + · have : Infinite J := ⟨h⟩ + let _ := Encodable.ofCountable J + let _ := Denumerable.ofEncodableOfInfinite J + exact hasExactLimitsOfShape_of_initial C (Discrete.equivalence (Denumerable.eqv J)).inverse + +section EpiMono + +open Functor + +section + +variable [HasZeroMorphisms C] [HasFiniteBiproducts C] + +noncomputable instance hasExactColimitsOfShape_discrete_finite (J : Type*) [Finite J] : + HasExactColimitsOfShape (Discrete J) C where + preservesFiniteLimits := preservesFiniteLimits_of_natIso HasBiproductsOfShape.colimIsoLim.symm + +noncomputable instance hasExactLimitsOfShape_discrete_finite {J : Type*} [Finite J] : + HasExactLimitsOfShape (Discrete J) C where + preservesFiniteColimits := preservesFiniteColimits_of_natIso HasBiproductsOfShape.colimIsoLim + +/-- +Checking AB of shape `Discrete ℕ` is enough for countable AB4, provided that the category has +finite biproducts and finite limits. +-/ +lemma CountableAB4.of_hasExactColimitsOfShape_nat [HasFiniteLimits C] [HasCountableCoproducts C] + [HasExactColimitsOfShape (Discrete ℕ) C] : CountableAB4 C := by + apply (config := { allowSynthFailures := true }) + CountableAB4.of_hasExactColimitsOfShape_nat_and_finite + exact fun _ ↦ inferInstance + +/-- +Checking AB* of shape `Discrete ℕ` is enough for countable AB4*, provided that the category has +finite biproducts and finite colimits. +-/ +lemma CountableAB4Star.of_hasExactLimitsOfShape_nat [HasFiniteColimits C] + [HasCountableProducts C] [HasExactLimitsOfShape (Discrete ℕ) C] : CountableAB4Star C := by + apply (config := { allowSynthFailures := true }) + CountableAB4Star.of_hasExactLimitsOfShape_nat_and_finite + exact fun _ ↦ inferInstance + end +variable [Abelian C] (J : Type u') [Category.{v'} J] + +attribute [local instance] preservesBinaryBiproducts_of_preservesBinaryCoproducts + preservesBinaryBiproducts_of_preservesBinaryProducts + +/-- +If `colim` of shape `J` into an abelian category `C` preserves monomorphisms, then `C` has AB of +shape `J`. +-/ +lemma hasExactColimitsOfShape_of_preservesMono [HasColimitsOfShape J C] + [PreservesMonomorphisms (colim (J := J) (C := C))] : HasExactColimitsOfShape J C where + preservesFiniteLimits := by + apply (config := { allowSynthFailures := true }) preservesFiniteLimits_of_preservesHomology + · exact preservesHomology_of_preservesMonos_and_cokernels _ + · exact additive_of_preservesBinaryBiproducts _ + +/-- +If `lim` of shape `J` into an abelian category `C` preserves epimorphisms, then `C` has AB* of +shape `J`. +-/ +lemma hasExactLimitsOfShape_of_preservesEpi [HasLimitsOfShape J C] + [PreservesEpimorphisms (lim (J := J) (C := C))] : HasExactLimitsOfShape J C where + preservesFiniteColimits := by + apply (config := { allowSynthFailures := true }) preservesFiniteColimits_of_preservesHomology + · exact preservesHomology_of_preservesEpis_and_kernels _ + · exact additive_of_preservesBinaryBiproducts _ + +end EpiMono + end CategoryTheory diff --git a/Mathlib/CategoryTheory/Limits/Filtered.lean b/Mathlib/CategoryTheory/Limits/Filtered.lean index 536b215594402..f30d4ab31ca83 100644 --- a/Mathlib/CategoryTheory/Limits/Filtered.lean +++ b/Mathlib/CategoryTheory/Limits/Filtered.lean @@ -17,7 +17,7 @@ Furthermore, we define the type classes `HasCofilteredLimitsOfSize` and `HasFilt -/ -universe w' w v u +universe w' w w₂' w₂ v u noncomputable section @@ -103,6 +103,32 @@ instance (priority := 100) hasColimitsOfShape_of_has_filtered_colimits HasColimitsOfShape I C := HasFilteredColimitsOfSize.HasColimitsOfShape _ +lemma hasCofilteredLimitsOfSize_of_univLE [UnivLE.{w, w₂}] [UnivLE.{w', w₂'}] + [HasCofilteredLimitsOfSize.{w₂', w₂} C] : + HasCofilteredLimitsOfSize.{w', w} C where + HasLimitsOfShape J := + haveI := IsCofiltered.of_equivalence ((ShrinkHoms.equivalence.{w₂'} J).trans <| + Shrink.equivalence.{w₂} (ShrinkHoms.{w} J)) + hasLimitsOfShape_of_equivalence ((ShrinkHoms.equivalence.{w₂'} J).trans <| + Shrink.equivalence.{w₂} (ShrinkHoms.{w} J)).symm + +lemma hasCofilteredLimitsOfSize_shrink [HasCofilteredLimitsOfSize.{max w' w₂', max w w₂} C] : + HasCofilteredLimitsOfSize.{w', w} C := + hasCofilteredLimitsOfSize_of_univLE.{w', w, max w' w₂', max w w₂} + +lemma hasFilteredColimitsOfSize_of_univLE [UnivLE.{w, w₂}] [UnivLE.{w', w₂'}] + [HasFilteredColimitsOfSize.{w₂', w₂} C] : + HasFilteredColimitsOfSize.{w', w} C where + HasColimitsOfShape J := + haveI := IsFiltered.of_equivalence ((ShrinkHoms.equivalence.{w₂'} J).trans <| + Shrink.equivalence.{w₂} (ShrinkHoms.{w} J)) + hasColimitsOfShape_of_equivalence ((ShrinkHoms.equivalence.{w₂'} J).trans <| + Shrink.equivalence.{w₂} (ShrinkHoms.{w} J)).symm + +lemma hasFilteredColimitsOfSize_shrink [HasFilteredColimitsOfSize.{max w' w₂', max w w₂} C] : + HasFilteredColimitsOfSize.{w', w} C := + hasFilteredColimitsOfSize_of_univLE.{w', w, max w' w₂', max w w₂} + end Limits end CategoryTheory diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Products.lean b/Mathlib/CategoryTheory/Limits/Shapes/Products.lean index ed9bd8a77cbda..02cc099007693 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Products.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Products.lean @@ -653,12 +653,17 @@ abbrev HasCoproducts := variable {C} -theorem has_smallest_products_of_hasProducts [HasProducts.{w} C] : HasProducts.{0} C := fun J => - hasLimitsOfShape_of_equivalence (Discrete.equivalence Equiv.ulift : Discrete (ULift.{w} J) ≌ _) +lemma hasProducts_shrink [HasProducts.{max w w'} C] : HasProducts.{w} C := fun J => + hasLimitsOfShape_of_equivalence (Discrete.equivalence Equiv.ulift : Discrete (ULift.{w'} J) ≌ _) + +lemma hasCoproducts_shrink [HasCoproducts.{max w w'} C] : HasCoproducts.{w} C := fun J => + hasColimitsOfShape_of_equivalence (Discrete.equivalence Equiv.ulift : Discrete (ULift.{w'} J) ≌ _) + +theorem has_smallest_products_of_hasProducts [HasProducts.{w} C] : HasProducts.{0} C := + hasProducts_shrink theorem has_smallest_coproducts_of_hasCoproducts [HasCoproducts.{w} C] : HasCoproducts.{0} C := - fun J => - hasColimitsOfShape_of_equivalence (Discrete.equivalence Equiv.ulift : Discrete (ULift.{w} J) ≌ _) + hasCoproducts_shrink theorem hasProducts_of_limit_fans (lf : ∀ {J : Type w} (f : J → C), Fan f) (lf_isLimit : ∀ {J : Type w} (f : J → C), IsLimit (lf f)) : HasProducts.{w} C := From 574c5f1bfe06a4bf0e9d1c4e86b754121a7f0d0b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Fri, 6 Dec 2024 15:40:36 +0000 Subject: [PATCH 552/829] feat(MeasureTheory): Stieltjes measure of various intervals (#19625) From the TestingLowerBounds project. Co-authored-by: Lorenzo Luccioli Co-authored-by: RemyDegenne --- Mathlib/MeasureTheory/Measure/Stieltjes.lean | 54 +++++++++++++++++++- 1 file changed, 53 insertions(+), 1 deletion(-) diff --git a/Mathlib/MeasureTheory/Measure/Stieltjes.lean b/Mathlib/MeasureTheory/Measure/Stieltjes.lean index 262bd709454e1..a381f2f16d3fb 100644 --- a/Mathlib/MeasureTheory/Measure/Stieltjes.lean +++ b/Mathlib/MeasureTheory/Measure/Stieltjes.lean @@ -129,6 +129,8 @@ instance : Module ℝ≥0 StieltjesFunction where add_smul _ _ _ := ext fun _ ↦ add_mul _ _ _ zero_smul _ := ext fun _ ↦ zero_mul _ +@[simp] lemma zero_apply (x : ℝ) : (0 : StieltjesFunction) x = 0 := rfl + @[simp] lemma add_apply (f g : StieltjesFunction) (x : ℝ) : (f + g) x = f x + g x := rfl /-- If a function `f : ℝ → ℝ` is monotone, then the function mapping `x` to the right limit of `f` @@ -430,6 +432,12 @@ theorem measure_Iic {l : ℝ} (hf : Tendsto f atBot (𝓝 l)) (x : ℝ) : simp_rw [measure_Ioc] exact ENNReal.tendsto_ofReal (Tendsto.const_sub _ hf) +lemma measure_Iio {l : ℝ} (hf : Tendsto f atBot (𝓝 l)) (x : ℝ) : + f.measure (Iio x) = ofReal (leftLim f x - l) := by + rw [← Iic_diff_right, measure_diff _ (nullMeasurableSet_singleton x), measure_singleton, + f.measure_Iic hf, ← ofReal_sub _ (sub_nonneg.mpr <| Monotone.leftLim_le f.mono' le_rfl)] + <;> simp + theorem measure_Ici {l : ℝ} (hf : Tendsto f atTop (𝓝 l)) (x : ℝ) : f.measure (Ici x) = ofReal (l - leftLim f x) := by refine tendsto_nhds_unique (tendsto_measure_Ico_atTop _ _) ?_ @@ -441,12 +449,52 @@ theorem measure_Ici {l : ℝ} (hf : Tendsto f atTop (𝓝 l)) (x : ℝ) : rw [tendsto_atTop_atTop] exact fun y => ⟨y + 1, fun z hyz => by rwa [le_sub_iff_add_le]⟩ +lemma measure_Ioi {l : ℝ} (hf : Tendsto f atTop (𝓝 l)) (x : ℝ) : + f.measure (Ioi x) = ofReal (l - f x) := by + rw [← Ici_diff_left, measure_diff _ (nullMeasurableSet_singleton x), measure_singleton, + f.measure_Ici hf, ← ofReal_sub _ (sub_nonneg.mpr <| Monotone.leftLim_le f.mono' le_rfl)] + <;> simp + +lemma measure_Ioi_of_tendsto_atTop_atTop (hf : Tendsto f atTop atTop) (x : ℝ) : + f.measure (Ioi x) = ∞ := by + refine ENNReal.eq_top_of_forall_nnreal_le fun r ↦ ?_ + obtain ⟨N, hN⟩ := eventually_atTop.mp (tendsto_atTop.mp hf (r + f x)) + exact (f.measure_Ioc x (max x N) ▸ ENNReal.coe_nnreal_eq r ▸ (ENNReal.ofReal_le_ofReal <| + le_tsub_of_add_le_right <| hN _ (le_max_right x N))).trans (measure_mono Ioc_subset_Ioi_self) + +lemma measure_Ici_of_tendsto_atTop_atTop (hf : Tendsto f atTop atTop) (x : ℝ) : + f.measure (Ici x) = ∞ := by + rw [← top_le_iff, ← f.measure_Ioi_of_tendsto_atTop_atTop hf x] + exact measure_mono Ioi_subset_Ici_self + +lemma measure_Iic_of_tendsto_atBot_atBot (hf : Tendsto f atBot atBot) (x : ℝ) : + f.measure (Iic x) = ∞ := by + refine ENNReal.eq_top_of_forall_nnreal_le fun r ↦ ?_ + obtain ⟨N, hN⟩ := eventually_atBot.mp (tendsto_atBot.mp hf (f x - r)) + exact (f.measure_Ioc (min x N) x ▸ ENNReal.coe_nnreal_eq r ▸ (ENNReal.ofReal_le_ofReal <| + le_sub_comm.mp <| hN _ (min_le_right x N))).trans (measure_mono Ioc_subset_Iic_self) + +lemma measure_Iio_of_tendsto_atBot_atBot (hf : Tendsto f atBot atBot) (x : ℝ) : + f.measure (Iio x) = ∞ := by + rw [← top_le_iff, ← f.measure_Iic_of_tendsto_atBot_atBot hf (x - 1)] + exact measure_mono <| Set.Iic_subset_Iio.mpr <| sub_one_lt x + theorem measure_univ {l u : ℝ} (hfl : Tendsto f atBot (𝓝 l)) (hfu : Tendsto f atTop (𝓝 u)) : f.measure univ = ofReal (u - l) := by refine tendsto_nhds_unique (tendsto_measure_Iic_atTop _) ?_ simp_rw [measure_Iic f hfl] exact ENNReal.tendsto_ofReal (Tendsto.sub_const hfu _) +lemma measure_univ_of_tendsto_atTop_atTop (hf : Tendsto f atTop atTop) : + f.measure univ = ∞ := by + rw [← top_le_iff, ← f.measure_Ioi_of_tendsto_atTop_atTop hf 0] + exact measure_mono (subset_univ _) + +lemma measure_univ_of_tendsto_atBot_atBot (hf : Tendsto f atBot atBot) : + f.measure univ = ∞ := by + rw [← top_le_iff, ← f.measure_Iio_of_tendsto_atBot_atBot hf 0] + exact measure_mono (subset_univ _) + lemma isFiniteMeasure {l u : ℝ} (hfl : Tendsto f atBot (𝓝 l)) (hfu : Tendsto f atTop (𝓝 u)) : IsFiniteMeasure f.measure := ⟨by simp [f.measure_univ hfl hfu]⟩ @@ -490,9 +538,13 @@ lemma eq_of_measure_of_eq (g : StieltjesFunction) {y : ℝ} · rw [sub_nonneg] exact f.mono hxy +@[simp] +lemma measure_zero : StieltjesFunction.measure 0 = 0 := + Measure.ext_of_Ioc _ _ (by simp) + @[simp] lemma measure_const (c : ℝ) : (StieltjesFunction.const c).measure = 0 := - Measure.ext_of_Ioc _ _ (fun _ _ _ ↦ by simp) + Measure.ext_of_Ioc _ _ (by simp) @[simp] lemma measure_add (f g : StieltjesFunction) : (f + g).measure = f.measure + g.measure := by From 6d1d83d97a6a3eb196b5db29f287f6716f233f1e Mon Sep 17 00:00:00 2001 From: Arend Mellendijk Date: Fri, 6 Dec 2024 16:19:18 +0000 Subject: [PATCH 553/829] feat: Count the number of pairs of `Nat`s whose `lcm` is `n` (#17614) This result is central to controlling the error term in my formalisation of the Selberg sieve. Specifically to count the number of terms in a sum. Co-authored-by: Eric Wieser Co-authored-by: Arend Mellendijk --- Mathlib/Algebra/Order/Antidiag/Nat.lean | 80 ++++++++++++++++++++++++- 1 file changed, 78 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Order/Antidiag/Nat.lean b/Mathlib/Algebra/Order/Antidiag/Nat.lean index 40c8db2fe9e60..88a8cc741a284 100644 --- a/Mathlib/Algebra/Order/Antidiag/Nat.lean +++ b/Mathlib/Algebra/Order/Antidiag/Nat.lean @@ -11,11 +11,13 @@ import Mathlib.NumberTheory.ArithmeticFunction # Sets of tuples with a fixed product This file defines the finite set of `d`-tuples of natural numbers with a fixed product `n` as -`Nat.finMulAntidiagonal`. +`Nat.finMulAntidiag`. ## Main Results * There are `d^(ω n)` ways to write `n` as a product of `d` natural numbers, when `n` is squarefree -(`card_finMulAntidiagonal_of_squarefree`) +(`card_finMulAntidiag_of_squarefree`) +* There are `3^(ω n)` pairs of natural numbers whose `lcm` is `n`, when `n` is squarefree +(`card_pair_lcm_eq`) -/ open Finset @@ -239,4 +241,78 @@ theorem card_finMulAntidiag_of_squarefree {d n : ℕ} (hn : Squarefree n) : ArithmeticFunction.cardDistinctFactors_apply, ← List.card_toFinset, toFinset_factors, Finset.card_fin] +theorem finMulAntidiag_three {n : ℕ} : + ∀ a ∈ finMulAntidiag 3 n, a 0 * a 1 * a 2 = n := by + intro a ha + rw [← (mem_finMulAntidiag.mp ha).1, Fin.prod_univ_three a] + +namespace card_pair_lcm_eq + +/-! +The following private declarations are ingredients for the proof of `card_pair_lcm_eq`. +-/ + +@[reducible] +private def f {n : ℕ} : ∀ a ∈ finMulAntidiag 3 n, ℕ × ℕ := fun a _ => (a 0 * a 1, a 0 * a 2) + +private theorem f_img {n : ℕ} (hn : Squarefree n) (a : Fin 3 → ℕ) + (ha : a ∈ finMulAntidiag 3 n) : + f a ha ∈ Finset.filter (fun ⟨x, y⟩ => x.lcm y = n) (n.divisors ×ˢ n.divisors) := by + rw [mem_filter, Finset.mem_product, mem_divisors, mem_divisors] + refine ⟨⟨⟨?_, hn.ne_zero⟩, ⟨?_, hn.ne_zero⟩⟩, ?_⟩ <;> rw [f, ← finMulAntidiag_three a ha] + · apply dvd_mul_right + · use a 1; ring + dsimp only + rw [lcm_mul_left, Nat.Coprime.lcm_eq_mul] + · ring + refine coprime_of_squarefree_mul (hn.squarefree_of_dvd ?_) + use a 0; rw [← finMulAntidiag_three a ha]; ring + +private theorem f_inj {n : ℕ} (a : Fin 3 → ℕ) (ha : a ∈ finMulAntidiag 3 n) + (b : Fin 3 → ℕ) (hb : b ∈ finMulAntidiag 3 n) (hfab : f a ha = f b hb) : + a = b := by + obtain ⟨hfab1, hfab2⟩ := Prod.mk.inj hfab + have hprods : a 0 * a 1 * a 2 = a 0 * a 1 * b 2 := by + rw [finMulAntidiag_three a ha, hfab1, finMulAntidiag_three b hb] + have hab2 : a 2 = b 2 := by + rw [← mul_right_inj' <| mul_ne_zero (ne_zero_of_mem_finMulAntidiag ha 0) + (ne_zero_of_mem_finMulAntidiag ha 1)] + exact hprods + have hab0 : a 0 = b 0 := by + rw [hab2] at hfab2 + exact (mul_left_inj' <| ne_zero_of_mem_finMulAntidiag hb 2).mp hfab2; + have hab1 : a 1 = b 1 := by + rw [hab0] at hfab1 + exact (mul_right_inj' <| ne_zero_of_mem_finMulAntidiag hb 0).mp hfab1; + funext i; fin_cases i <;> assumption + +private theorem f_surj {n : ℕ} (hn : n ≠ 0) (b : ℕ × ℕ) + (hb : b ∈ Finset.filter (fun ⟨x, y⟩ => x.lcm y = n) (n.divisors ×ˢ n.divisors)) : + ∃ (a : Fin 3 → ℕ) (ha : a ∈ finMulAntidiag 3 n), f a ha = b := by + dsimp only at hb + let g := b.fst.gcd b.snd + let a := ![g, b.fst/g, b.snd/g] + have ha : a ∈ finMulAntidiag 3 n := by + rw [mem_finMulAntidiag] + rw [mem_filter, Finset.mem_product] at hb + refine ⟨?_, hn⟩ + · rw [Fin.prod_univ_three a] + simp only [a, Matrix.cons_val_zero, Matrix.cons_val_one, Matrix.head_cons, + Matrix.cons_val_two, Matrix.tail_cons] + rw [Nat.mul_div_cancel_left' (Nat.gcd_dvd_left _ _), ← hb.2, lcm, + Nat.mul_div_assoc b.fst (Nat.gcd_dvd_right b.fst b.snd)] + use a; use ha + apply Prod.ext <;> simp only [Matrix.cons_val_zero, Matrix.cons_val_one, Matrix.head_cons] + <;> apply Nat.mul_div_cancel' + · apply Nat.gcd_dvd_left + · apply Nat.gcd_dvd_right + +end card_pair_lcm_eq + +open card_pair_lcm_eq in +theorem card_pair_lcm_eq {n : ℕ} (hn : Squarefree n) : + #{p ∈ (n.divisors ×ˢ n.divisors) | p.1.lcm p.2 = n} = 3 ^ ω n := by + rw [← card_finMulAntidiag_of_squarefree hn, eq_comm] + apply Finset.card_bij f (f_img hn) (f_inj) (f_surj hn.ne_zero) + end Nat From a4bbd290638a8b8ef8896655cf1fa98fba66a337 Mon Sep 17 00:00:00 2001 From: Mitchell Lee Date: Fri, 6 Dec 2024 16:45:26 +0000 Subject: [PATCH 554/829] feat: specialize theorems about filters disjoint from the cocompact filter to singleton sets (#19717) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Let `X, Y` be types and let `f` be a filter on `Y` which is disjoint from the cocompact filter. We prove that `𝓝 x ×ˢ f ≤ 𝓝ˢ ({x} ×ˢ Set.univ)` for all `x : X`. This is a special case of the existing theorem `nhdsSet_prod_le_of_disjoint_cocompact`. We also prove the analogous result with `X` and `Y` swapped. --- Mathlib/Topology/Compactness/Compact.lean | 12 +++++++++++- 1 file changed, 11 insertions(+), 1 deletion(-) diff --git a/Mathlib/Topology/Compactness/Compact.lean b/Mathlib/Topology/Compactness/Compact.lean index a28aac3af3af1..15e1b303bf533 100644 --- a/Mathlib/Topology/Compactness/Compact.lean +++ b/Mathlib/Topology/Compactness/Compact.lean @@ -698,7 +698,7 @@ theorem nhdsSet_prod_le_of_disjoint_cocompact {f : Filter Y} (hs : IsCompact s) _ = 𝓝ˢ (s ×ˢ K) := (hs.nhdsSet_prod_eq hK).symm _ ≤ 𝓝ˢ (s ×ˢ Set.univ) := nhdsSet_mono (prod_mono_right le_top) -theorem prod_nhdsSet_le_of_disjoint_cocompact {f : Filter X} (ht : IsCompact t) +theorem prod_nhdsSet_le_of_disjoint_cocompact {t : Set Y} {f : Filter X} (ht : IsCompact t) (hf : Disjoint f (Filter.cocompact X)) : f ×ˢ 𝓝ˢ t ≤ 𝓝ˢ (Set.univ ×ˢ t) := by obtain ⟨K, hKf, hK⟩ := (disjoint_cocompact_right f).mp hf @@ -709,6 +709,16 @@ theorem prod_nhdsSet_le_of_disjoint_cocompact {f : Filter X} (ht : IsCompact t) _ = 𝓝ˢ (K ×ˢ t) := (hK.nhdsSet_prod_eq ht).symm _ ≤ 𝓝ˢ (Set.univ ×ˢ t) := nhdsSet_mono (prod_mono_left le_top) +theorem nhds_prod_le_of_disjoint_cocompact {f : Filter Y} (x : X) + (hf : Disjoint f (Filter.cocompact Y)) : + 𝓝 x ×ˢ f ≤ 𝓝ˢ ({x} ×ˢ Set.univ) := by + simpa using nhdsSet_prod_le_of_disjoint_cocompact isCompact_singleton hf + +theorem prod_nhds_le_of_disjoint_cocompact {f : Filter X} (y : Y) + (hf : Disjoint f (Filter.cocompact X)) : + f ×ˢ 𝓝 y ≤ 𝓝ˢ (Set.univ ×ˢ {y}) := by + simpa using prod_nhdsSet_le_of_disjoint_cocompact isCompact_singleton hf + /-- If `s` and `t` are compact sets and `n` is an open neighborhood of `s × t`, then there exist open neighborhoods `u ⊇ s` and `v ⊇ t` such that `u × v ⊆ n`. From 75b9e9130c8c5495259e4f04622fb4575bae0f11 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Fri, 6 Dec 2024 17:36:29 +0000 Subject: [PATCH 555/829] feat(Algebra/Order/SuccPred): `IsSuccLimit` theorems for `SuccAddOrder` (#19060) --- Mathlib/Algebra/Order/SuccPred.lean | 42 ++++++++++++++++++++++++++++- Mathlib/Order/SuccPred/Limit.lean | 3 +++ 2 files changed, 44 insertions(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Order/SuccPred.lean b/Mathlib/Algebra/Order/SuccPred.lean index 185ea8250d6f9..faf2e279f4aa8 100644 --- a/Mathlib/Algebra/Order/SuccPred.lean +++ b/Mathlib/Algebra/Order/SuccPred.lean @@ -6,7 +6,7 @@ Authors: Violeta Hernández Palacios, Yaël Dillies import Mathlib.Algebra.Group.Basic import Mathlib.Algebra.Order.ZeroLEOne import Mathlib.Data.Int.Cast.Defs -import Mathlib.Order.SuccPred.Archimedean +import Mathlib.Order.SuccPred.Limit /-! # Interaction between successors and arithmetic @@ -138,6 +138,46 @@ theorem covBy_iff_sub_one_eq [Sub α] [One α] [PredSubOrder α] [NoMinOrder α] rw [← pred_eq_sub_one] exact pred_eq_iff_covBy.symm +theorem IsSuccPrelimit.add_one_lt [Add α] [One α] [SuccAddOrder α] + (hx : IsSuccPrelimit x) (hy : y < x) : y + 1 < x := by + rw [← succ_eq_add_one] + exact hx.succ_lt hy + +theorem IsPredPrelimit.lt_sub_one [Sub α] [One α] [PredSubOrder α] + (hx : IsPredPrelimit x) (hy : x < y) : x < y - 1 := by + rw [← pred_eq_sub_one] + exact hx.lt_pred hy + +theorem IsSuccLimit.add_one_lt [Add α] [One α] [SuccAddOrder α] + (hx : IsSuccLimit x) (hy : y < x) : y + 1 < x := + hx.isSuccPrelimit.add_one_lt hy + +theorem IsPredLimit.lt_sub_one [Sub α] [One α] [PredSubOrder α] + (hx : IsPredLimit x) (hy : x < y) : x < y - 1 := + hx.isPredPrelimit.lt_sub_one hy + +theorem IsSuccPrelimit.add_natCast_lt [AddMonoidWithOne α] [SuccAddOrder α] + (hx : IsSuccPrelimit x) (hy : y < x) : ∀ n : ℕ, y + n < x + | 0 => by simpa + | n + 1 => by + rw [Nat.cast_add_one, ← add_assoc] + exact hx.add_one_lt (hx.add_natCast_lt hy n) + +theorem IsPredPrelimit.lt_sub_natCast [AddCommGroupWithOne α] [PredSubOrder α] + (hx : IsPredPrelimit x) (hy : x < y) : ∀ n : ℕ, x < y - n + | 0 => by simpa + | n + 1 => by + rw [Nat.cast_add_one, ← sub_sub] + exact hx.lt_sub_one (hx.lt_sub_natCast hy n) + +theorem IsSuccLimit.add_natCast_lt [AddMonoidWithOne α] [SuccAddOrder α] + (hx : IsSuccLimit x) (hy : y < x) : ∀ n : ℕ, y + n < x := + hx.isSuccPrelimit.add_natCast_lt hy + +theorem IsPredLimit.lt_sub_natCast [AddCommGroupWithOne α] [PredSubOrder α] + (hx : IsPredLimit x) (hy : x < y) : ∀ n : ℕ, x < y - n := + hx.isPredPrelimit.lt_sub_natCast hy + end PartialOrder section LinearOrder diff --git a/Mathlib/Order/SuccPred/Limit.lean b/Mathlib/Order/SuccPred/Limit.lean index e23e9f3184fd3..6170a1e80ff57 100644 --- a/Mathlib/Order/SuccPred/Limit.lean +++ b/Mathlib/Order/SuccPred/Limit.lean @@ -499,6 +499,9 @@ variable [PartialOrder α] theorem isPredLimit_iff [OrderTop α] : IsPredLimit a ↔ a ≠ ⊤ ∧ IsPredPrelimit a := by rw [IsPredLimit, isMax_iff_eq_top] +theorem IsPredLimit.lt_top [OrderTop α] (h : IsPredLimit a) : a < ⊤ := + h.ne_top.lt_top + variable [PredOrder α] theorem isPredPrelimit_of_pred_ne (h : ∀ b, pred b ≠ a) : IsPredPrelimit a := fun b hba => From 5f596a441e387004ca7855fb1b91bceeecd710d5 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Fri, 6 Dec 2024 18:07:38 +0000 Subject: [PATCH 556/829] feat(Data/Fin/VecNotation): remove `simp` from `vecCons_const` (#19457) There are a couple of simp lemmas that trigger on `vecCons`, but not on the reduced form, hence let's remove the `simp` from `vecCons_const`. Co-authored-by: Moritz Firsching --- Mathlib/Data/Fin/VecNotation.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Mathlib/Data/Fin/VecNotation.lean b/Mathlib/Data/Fin/VecNotation.lean index 1b8ca2379a123..aeb31bb6c4a94 100644 --- a/Mathlib/Data/Fin/VecNotation.lean +++ b/Mathlib/Data/Fin/VecNotation.lean @@ -167,7 +167,6 @@ theorem range_cons_cons_empty (x y : α) (u : Fin 0 → α) : Set.range (vecCons x <| vecCons y u) = {x, y} := by rw [range_cons, range_cons_empty, Set.singleton_union] -@[simp] theorem vecCons_const (a : α) : (vecCons a fun _ : Fin n => a) = fun _ => a := funext <| Fin.forall_iff_succ.2 ⟨rfl, cons_val_succ _ _⟩ From dd148f49c40b0fd36156ee4bb246e4e2d42b05cf Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 6 Dec 2024 18:07:40 +0000 Subject: [PATCH 557/829] chore: minor cleanup of Cardinal.Basic (#19631) Just rearranging material within the file. Not splitting today. --- Mathlib/SetTheory/Cardinal/Basic.lean | 130 +++++++++++++------------- 1 file changed, 67 insertions(+), 63 deletions(-) diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index cb32fdc0b2c53..12ba9cbde683e 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -177,54 +177,6 @@ def map₂ (f : Type u → Type v → Type w) (hf : ∀ α β γ δ, α ≃ β Cardinal.{u} → Cardinal.{v} → Cardinal.{w} := Quotient.map₂ f fun α β ⟨e₁⟩ γ δ ⟨e₂⟩ => ⟨hf α β γ δ e₁ e₂⟩ -/-- We define the order on cardinal numbers by `#α ≤ #β` if and only if - there exists an embedding (injective function) from α to β. -/ -instance : LE Cardinal.{u} := - ⟨fun q₁ q₂ => - Quotient.liftOn₂ q₁ q₂ (fun α β => Nonempty <| α ↪ β) fun _ _ _ _ ⟨e₁⟩ ⟨e₂⟩ => - propext ⟨fun ⟨e⟩ => ⟨e.congr e₁ e₂⟩, fun ⟨e⟩ => ⟨e.congr e₁.symm e₂.symm⟩⟩⟩ - -instance partialOrder : PartialOrder Cardinal.{u} where - le := (· ≤ ·) - le_refl := by - rintro ⟨α⟩ - exact ⟨Embedding.refl _⟩ - le_trans := by - rintro ⟨α⟩ ⟨β⟩ ⟨γ⟩ ⟨e₁⟩ ⟨e₂⟩ - exact ⟨e₁.trans e₂⟩ - le_antisymm := by - rintro ⟨α⟩ ⟨β⟩ ⟨e₁⟩ ⟨e₂⟩ - exact Quotient.sound (e₁.antisymm e₂) - -instance linearOrder : LinearOrder Cardinal.{u} := - { Cardinal.partialOrder with - le_total := by - rintro ⟨α⟩ ⟨β⟩ - apply Embedding.total - decidableLE := Classical.decRel _ } - -theorem le_def (α β : Type u) : #α ≤ #β ↔ Nonempty (α ↪ β) := - Iff.rfl - -theorem mk_le_of_injective {α β : Type u} {f : α → β} (hf : Injective f) : #α ≤ #β := - ⟨⟨f, hf⟩⟩ - -theorem _root_.Function.Embedding.cardinal_le {α β : Type u} (f : α ↪ β) : #α ≤ #β := - ⟨f⟩ - -theorem mk_le_of_surjective {α β : Type u} {f : α → β} (hf : Surjective f) : #β ≤ #α := - ⟨Embedding.ofSurjective f hf⟩ - -theorem le_mk_iff_exists_set {c : Cardinal} {α : Type u} : c ≤ #α ↔ ∃ p : Set α, #p = c := - ⟨inductionOn c fun _ ⟨⟨f, hf⟩⟩ => ⟨Set.range f, (Equiv.ofInjective f hf).cardinal_eq.symm⟩, - fun ⟨_, e⟩ => e ▸ ⟨⟨Subtype.val, fun _ _ => Subtype.eq⟩⟩⟩ - -theorem mk_subtype_le {α : Type u} (p : α → Prop) : #(Subtype p) ≤ #α := - ⟨Embedding.subtype p⟩ - -theorem mk_set_le (s : Set α) : #s ≤ #α := - mk_subtype_le s - /-! ### Lifting cardinals to a higher universe -/ /-- The universe lift operation on cardinals. You can specify the universes explicitly with @@ -281,21 +233,6 @@ lemma mk_preimage_down {s : Set α} : #(ULift.down.{v} ⁻¹' s) = lift.{v} (#s) ULift.up_bijective.comp (restrictPreimage_bijective _ (ULift.down_bijective)) exact Equiv.ofBijective f this -theorem out_embedding {c c' : Cardinal} : c ≤ c' ↔ Nonempty (c.out ↪ c'.out) := by - conv_lhs => rw [← Cardinal.mk_out c, ← Cardinal.mk_out c', le_def] - -theorem lift_mk_le {α : Type v} {β : Type w} : - lift.{max u w} #α ≤ lift.{max u v} #β ↔ Nonempty (α ↪ β) := - ⟨fun ⟨f⟩ => ⟨Embedding.congr Equiv.ulift Equiv.ulift f⟩, fun ⟨f⟩ => - ⟨Embedding.congr Equiv.ulift.symm Equiv.ulift.symm f⟩⟩ - -/-- A variant of `Cardinal.lift_mk_le` with specialized universes. -Because Lean often can not realize it should use this specialization itself, -we provide this statement separately so you don't have to solve the specialization problem either. --/ -theorem lift_mk_le' {α : Type u} {β : Type v} : lift.{v} #α ≤ lift.{u} #β ↔ Nonempty (α ↪ β) := - lift_mk_le.{0} - theorem lift_mk_eq {α : Type u} {β : Type v} : lift.{max v w} #α = lift.{max u w} #β ↔ Nonempty (α ≃ β) := Quotient.eq'.trans @@ -325,6 +262,73 @@ theorem lift_mk_shrink'' (α : Type max u v) [Small.{v} α] : Cardinal.lift.{u} #(Shrink.{v} α) = #α := by rw [← lift_umax, lift_mk_shrink.{max u v, v, 0} α, ← lift_umax, lift_id] +/-! ### Order on cardinals -/ + +/-- We define the order on cardinal numbers by `#α ≤ #β` if and only if + there exists an embedding (injective function) from α to β. -/ +instance : LE Cardinal.{u} := + ⟨fun q₁ q₂ => + Quotient.liftOn₂ q₁ q₂ (fun α β => Nonempty <| α ↪ β) fun _ _ _ _ ⟨e₁⟩ ⟨e₂⟩ => + propext ⟨fun ⟨e⟩ => ⟨e.congr e₁ e₂⟩, fun ⟨e⟩ => ⟨e.congr e₁.symm e₂.symm⟩⟩⟩ + +instance partialOrder : PartialOrder Cardinal.{u} where + le := (· ≤ ·) + le_refl := by + rintro ⟨α⟩ + exact ⟨Embedding.refl _⟩ + le_trans := by + rintro ⟨α⟩ ⟨β⟩ ⟨γ⟩ ⟨e₁⟩ ⟨e₂⟩ + exact ⟨e₁.trans e₂⟩ + le_antisymm := by + rintro ⟨α⟩ ⟨β⟩ ⟨e₁⟩ ⟨e₂⟩ + exact Quotient.sound (e₁.antisymm e₂) + +instance linearOrder : LinearOrder Cardinal.{u} := + { Cardinal.partialOrder with + le_total := by + rintro ⟨α⟩ ⟨β⟩ + apply Embedding.total + decidableLE := Classical.decRel _ } + +theorem le_def (α β : Type u) : #α ≤ #β ↔ Nonempty (α ↪ β) := + Iff.rfl + +theorem mk_le_of_injective {α β : Type u} {f : α → β} (hf : Injective f) : #α ≤ #β := + ⟨⟨f, hf⟩⟩ + +theorem _root_.Function.Embedding.cardinal_le {α β : Type u} (f : α ↪ β) : #α ≤ #β := + ⟨f⟩ + +theorem mk_le_of_surjective {α β : Type u} {f : α → β} (hf : Surjective f) : #β ≤ #α := + ⟨Embedding.ofSurjective f hf⟩ + +theorem le_mk_iff_exists_set {c : Cardinal} {α : Type u} : c ≤ #α ↔ ∃ p : Set α, #p = c := + ⟨inductionOn c fun _ ⟨⟨f, hf⟩⟩ => ⟨Set.range f, (Equiv.ofInjective f hf).cardinal_eq.symm⟩, + fun ⟨_, e⟩ => e ▸ ⟨⟨Subtype.val, fun _ _ => Subtype.eq⟩⟩⟩ + +theorem mk_subtype_le {α : Type u} (p : α → Prop) : #(Subtype p) ≤ #α := + ⟨Embedding.subtype p⟩ + +theorem mk_set_le (s : Set α) : #s ≤ #α := + mk_subtype_le s + +theorem out_embedding {c c' : Cardinal} : c ≤ c' ↔ Nonempty (c.out ↪ c'.out) := by + conv_lhs => rw [← Cardinal.mk_out c, ← Cardinal.mk_out c', le_def] + +theorem lift_mk_le {α : Type v} {β : Type w} : + lift.{max u w} #α ≤ lift.{max u v} #β ↔ Nonempty (α ↪ β) := + ⟨fun ⟨f⟩ => ⟨Embedding.congr Equiv.ulift Equiv.ulift f⟩, fun ⟨f⟩ => + ⟨Embedding.congr Equiv.ulift.symm Equiv.ulift.symm f⟩⟩ + +/-- A variant of `Cardinal.lift_mk_le` with specialized universes. +Because Lean often can not realize it should use this specialization itself, +we provide this statement separately so you don't have to solve the specialization problem either. +-/ +theorem lift_mk_le' {α : Type u} {β : Type v} : lift.{v} #α ≤ lift.{u} #β ↔ Nonempty (α ↪ β) := + lift_mk_le.{0} + +/-! ### `lift` sends `Cardinal.{u}` to an initial segment of `Cardinal.{max u v}`. -/ + /-- `Cardinal.lift` as an `InitialSeg`. -/ @[simps!] def liftInitialSeg : Cardinal.{u} ≤i Cardinal.{max u v} := by From e95e196cb074765317ad665f5c1570f1ef13dac6 Mon Sep 17 00:00:00 2001 From: Julian Berman Date: Fri, 6 Dec 2024 18:07:41 +0000 Subject: [PATCH 558/829] chore(RingTheory/Multiplicity): squeeze an import (#19764) This was importing linarith, but it isn't ever used in the file, all that seems needed is finding the instance of Ring for the integers, used in just Int.emultiplicity_natAbs. --- Mathlib/RingTheory/Multiplicity.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/Multiplicity.lean b/Mathlib/RingTheory/Multiplicity.lean index c1dc9a1ff5b51..19c092dd8f1ed 100644 --- a/Mathlib/RingTheory/Multiplicity.lean +++ b/Mathlib/RingTheory/Multiplicity.lean @@ -6,8 +6,8 @@ Authors: Robert Y. Lewis, Chris Hughes, Daniel Weber import Mathlib.Algebra.Associated.Basic import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Algebra.Ring.Divisibility.Basic +import Mathlib.Algebra.Ring.Int.Defs import Mathlib.Data.ENat.Basic -import Mathlib.Tactic.Linarith /-! # Multiplicity of a divisor From 29acab45301e724bfe34b5db6ca8793723abbf23 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 7 Dec 2024 03:05:19 +0000 Subject: [PATCH 559/829] feat(Pow/Real): compare `x^y` to `x` (#19720) --- .../Analysis/SpecialFunctions/Pow/Real.lean | 44 +++++++++++++++++-- 1 file changed, 40 insertions(+), 4 deletions(-) diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean index a0af1f9a0076e..0a7f64d0a15aa 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean @@ -698,14 +698,50 @@ theorem one_lt_rpow_iff (hx : 0 ≤ x) : 1 < x ^ y ↔ 1 < x ∧ 0 < y ∨ 0 < x · rcases _root_.em (y = 0) with (rfl | hy) <;> simp [*, lt_irrefl, (zero_lt_one' ℝ).not_lt] · simp [one_lt_rpow_iff_of_pos hx, hx] -theorem rpow_le_rpow_of_exponent_ge' (hx0 : 0 ≤ x) (hx1 : x ≤ 1) (hz : 0 ≤ z) (hyz : z ≤ y) : +/-- This is a more general but less convenient version of `rpow_le_rpow_of_exponent_ge`. +This version allows `x = 0`, so it explicitly forbids `x = y = 0`, `z ≠ 0`. -/ +theorem rpow_le_rpow_of_exponent_ge_of_imp (hx0 : 0 ≤ x) (hx1 : x ≤ 1) (hyz : z ≤ y) + (h : x = 0 → y = 0 → z = 0) : x ^ y ≤ x ^ z := by rcases eq_or_lt_of_le hx0 with (rfl | hx0') - · rcases eq_or_lt_of_le hz with (rfl | hz') - · exact (rpow_zero 0).symm ▸ rpow_le_one hx0 hx1 hyz - rw [zero_rpow, zero_rpow] <;> linarith + · rcases eq_or_ne y 0 with rfl | hy0 + · rw [h rfl rfl] + · rw [zero_rpow hy0] + apply zero_rpow_nonneg · exact rpow_le_rpow_of_exponent_ge hx0' hx1 hyz +/-- This version of `rpow_le_rpow_of_exponent_ge` allows `x = 0` but requires `0 ≤ z`. +See also `rpow_le_rpow_of_exponent_ge_of_imp` for the most general version. -/ +theorem rpow_le_rpow_of_exponent_ge' (hx0 : 0 ≤ x) (hx1 : x ≤ 1) (hz : 0 ≤ z) (hyz : z ≤ y) : + x ^ y ≤ x ^ z := + rpow_le_rpow_of_exponent_ge_of_imp hx0 hx1 hyz fun _ hy ↦ le_antisymm (hyz.trans_eq hy) hz + +theorem self_le_rpow_of_le_one (h₁ : 0 ≤ x) (h₂ : x ≤ 1) (h₃ : y ≤ 1) : x ≤ x ^ y := by + simpa only [rpow_one] + using rpow_le_rpow_of_exponent_ge_of_imp h₁ h₂ h₃ fun _ ↦ (absurd · one_ne_zero) + +theorem self_le_rpow_of_one_le (h₁ : 1 ≤ x) (h₂ : 1 ≤ y) : x ≤ x ^ y := by + simpa only [rpow_one] using rpow_le_rpow_of_exponent_le h₁ h₂ + +theorem rpow_le_self_of_le_one (h₁ : 0 ≤ x) (h₂ : x ≤ 1) (h₃ : 1 ≤ y) : x ^ y ≤ x := by + simpa only [rpow_one] + using rpow_le_rpow_of_exponent_ge_of_imp h₁ h₂ h₃ fun _ ↦ (absurd · (one_pos.trans_le h₃).ne') + +theorem rpow_le_self_of_one_le (h₁ : 1 ≤ x) (h₂ : y ≤ 1) : x ^ y ≤ x := by + simpa only [rpow_one] using rpow_le_rpow_of_exponent_le h₁ h₂ + +theorem self_lt_rpow_of_lt_one (h₁ : 0 < x) (h₂ : x < 1) (h₃ : y < 1) : x < x ^ y := by + simpa only [rpow_one] using rpow_lt_rpow_of_exponent_gt h₁ h₂ h₃ + +theorem self_lt_rpow_of_one_lt (h₁ : 1 < x) (h₂ : 1 < y) : x < x ^ y := by + simpa only [rpow_one] using rpow_lt_rpow_of_exponent_lt h₁ h₂ + +theorem rpow_lt_self_of_lt_one (h₁ : 0 < x) (h₂ : x < 1) (h₃ : 1 < y) : x ^ y < x := by + simpa only [rpow_one] using rpow_lt_rpow_of_exponent_gt h₁ h₂ h₃ + +theorem rpow_lt_self_of_one_lt (h₁ : 1 < x) (h₂ : y < 1) : x ^ y < x := by + simpa only [rpow_one] using rpow_lt_rpow_of_exponent_lt h₁ h₂ + theorem rpow_left_injOn {x : ℝ} (hx : x ≠ 0) : InjOn (fun y : ℝ => y ^ x) { y : ℝ | 0 ≤ y } := by rintro y hy z hz (hyz : y ^ x = z ^ x) rw [← rpow_one y, ← rpow_one z, ← mul_inv_cancel₀ hx, rpow_mul hy, rpow_mul hz, hyz] From cc6b7b9d6d4d64aee5e09e895d146dcf824c6350 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Sat, 7 Dec 2024 04:52:05 +0000 Subject: [PATCH 560/829] chore(Mathlib/NumberTheory/NumberField/*): rename completion -> Completion (#19646) We have `UniformSpace.Completion` so it should surely also be `AbsoluteValue.Completion v : Type*` and `NumberField.InfinitePlace.Completion v : Type*` (currently they are both `completion` and both types). Co-authored-by: Eric Wieser --- Mathlib/Analysis/Normed/Ring/WithAbs.lean | 22 ++--- .../NumberTheory/NumberField/AdeleRing.lean | 6 +- .../NumberTheory/NumberField/Completion.lean | 82 ++++++++++--------- 3 files changed, 57 insertions(+), 53 deletions(-) diff --git a/Mathlib/Analysis/Normed/Ring/WithAbs.lean b/Mathlib/Analysis/Normed/Ring/WithAbs.lean index d616c2ef77867..4cbcad5fc3e00 100644 --- a/Mathlib/Analysis/Normed/Ring/WithAbs.lean +++ b/Mathlib/Analysis/Normed/Ring/WithAbs.lean @@ -22,7 +22,7 @@ being used to define Archimedean completions of a number field. to assign and infer instances on a semiring that depend on absolute values. - `WithAbs.equiv v` : the canonical (type) equivalence between `WithAbs v` and `R`. - `WithAbs.ringEquiv v` : The canonical ring equivalence between `WithAbs v` and `R`. - - `AbsoluteValue.completion` : the uniform space completion of a field `K` according to the + - `AbsoluteValue.Completion` : the uniform space completion of a field `K` according to the uniform structure defined by the specified real absolute value. -/ @@ -141,18 +141,20 @@ open WithAbs variable {K : Type*} [Field K] (v : AbsoluteValue K ℝ) /-- The completion of a field with respect to a real absolute value. -/ -abbrev completion := UniformSpace.Completion (WithAbs v) +abbrev Completion := UniformSpace.Completion (WithAbs v) + +@[deprecated (since := "2024-12-01")] alias completion := Completion namespace Completion -instance : Coe K v.completion := +instance : Coe K v.Completion := inferInstanceAs <| Coe (WithAbs v) (UniformSpace.Completion (WithAbs v)) variable {L : Type*} [NormedField L] [CompleteSpace L] {f : WithAbs v →+* L} {v} /-- If the absolute value of a normed field factors through an embedding into another normed field -`L`, then we can extend that embedding to an embedding on the completion `v.completion →+* L`. -/ -abbrev extensionEmbedding_of_comp (h : ∀ x, ‖f x‖ = v x) : v.completion →+* L := +`L`, then we can extend that embedding to an embedding on the completion `v.Completion →+* L`. -/ +abbrev extensionEmbedding_of_comp (h : ∀ x, ‖f x‖ = v x) : v.Completion →+* L := UniformSpace.Completion.extensionHom _ (WithAbs.isUniformInducing_of_comp h).uniformContinuous.continuous @@ -162,8 +164,8 @@ theorem extensionEmbedding_of_comp_coe (h : ∀ x, ‖f x‖ = v x) (x : K) : (WithAbs.isUniformInducing_of_comp h).uniformContinuous.continuous] /-- If the absolute value of a normed field factors through an embedding into another normed field, -then the extended embedding `v.completion →+* L` preserves distances. -/ -theorem extensionEmbedding_dist_eq_of_comp (h : ∀ x, ‖f x‖ = v x) (x y : v.completion) : +then the extended embedding `v.Completion →+* L` preserves distances. -/ +theorem extensionEmbedding_dist_eq_of_comp (h : ∀ x, ‖f x‖ = v x) (x y : v.Completion) : dist (extensionEmbedding_of_comp h x) (extensionEmbedding_of_comp h y) = dist x y := by refine UniformSpace.Completion.induction_on₂ x y ?_ (fun x y => ?_) @@ -173,13 +175,13 @@ theorem extensionEmbedding_dist_eq_of_comp (h : ∀ x, ‖f x‖ = v x) (x y : v exact UniformSpace.Completion.dist_eq x y ▸ (WithAbs.isometry_of_comp h).dist_eq _ _ /-- If the absolute value of a normed field factors through an embedding into another normed field, -then the extended embedding `v.completion →+* L` is an isometry. -/ +then the extended embedding `v.Completion →+* L` is an isometry. -/ theorem isometry_extensionEmbedding_of_comp (h : ∀ x, ‖f x‖ = v x) : Isometry (extensionEmbedding_of_comp h) := Isometry.of_dist_eq <| extensionEmbedding_dist_eq_of_comp h /-- If the absolute value of a normed field factors through an embedding into another normed field, -then the extended embedding `v.completion →+* L` is a closed embedding. -/ +then the extended embedding `v.Completion →+* L` is a closed embedding. -/ theorem isClosedEmbedding_extensionEmbedding_of_comp (h : ∀ x, ‖f x‖ = v x) : IsClosedEmbedding (extensionEmbedding_of_comp h) := (isometry_extensionEmbedding_of_comp h).isClosedEmbedding @@ -187,7 +189,7 @@ theorem isClosedEmbedding_extensionEmbedding_of_comp (h : ∀ x, ‖f x‖ = v x /-- If the absolute value of a normed field factors through an embedding into another normed field that is locally compact, then the completion of the first normed field is also locally compact. -/ theorem locallyCompactSpace [LocallyCompactSpace L] (h : ∀ x, ‖f x‖ = v x) : - LocallyCompactSpace (v.completion) := + LocallyCompactSpace (v.Completion) := (isClosedEmbedding_extensionEmbedding_of_comp h).locallyCompactSpace end AbsoluteValue.Completion diff --git a/Mathlib/NumberTheory/NumberField/AdeleRing.lean b/Mathlib/NumberTheory/NumberField/AdeleRing.lean index 7c5add17d6901..fbf792b8edcbb 100644 --- a/Mathlib/NumberTheory/NumberField/AdeleRing.lean +++ b/Mathlib/NumberTheory/NumberField/AdeleRing.lean @@ -47,11 +47,11 @@ variable (K : Type*) [Field K] The infinite adele ring is the finite product of completions of a number field over its infinite places. See `NumberField.InfinitePlace` for the definition of an infinite place and -`NumberField.InfinitePlace.completion` for the associated completion. +`NumberField.InfinitePlace.Completion` for the associated completion. -/ /-- The infinite adele ring of a number field. -/ -def InfiniteAdeleRing := (v : InfinitePlace K) → v.completion +def InfiniteAdeleRing := (v : InfinitePlace K) → v.Completion namespace InfiniteAdeleRing @@ -82,7 +82,7 @@ abbrev ringEquiv_mixedSpace : InfiniteAdeleRing K ≃+* mixedEmbedding.mixedSpace K := RingEquiv.trans (RingEquiv.piEquivPiSubtypeProd (fun (v : InfinitePlace K) => IsReal v) - (fun (v : InfinitePlace K) => v.completion)) + (fun (v : InfinitePlace K) => v.Completion)) (RingEquiv.prodCongr (RingEquiv.piCongrRight (fun ⟨_, hv⟩ => Completion.ringEquiv_real_of_isReal hv)) (RingEquiv.trans diff --git a/Mathlib/NumberTheory/NumberField/Completion.lean b/Mathlib/NumberTheory/NumberField/Completion.lean index 901f318194033..155d6e66f7450 100644 --- a/Mathlib/NumberTheory/NumberField/Completion.lean +++ b/Mathlib/NumberTheory/NumberField/Completion.lean @@ -19,35 +19,35 @@ of instances is through the use of type synonyms. In this case, we use the type of a semiring. In particular this type synonym depends on an absolute value, which provides a systematic way of assigning and inferring instances of the semiring that also depend on an absolute value. The completion of a field at multiple absolute values is defined in -`Mathlib.Algebra.Ring.WithAbs` as `AbsoluteValue.completion`. The completion of a number +`Mathlib.Algebra.Ring.WithAbs` as `AbsoluteValue.Completion`. The completion of a number field at an infinite place is then derived in this file, as `InfinitePlace` is a subtype of `AbsoluteValue`. ## Main definitions - - `NumberField.InfinitePlace.completion` : the completion of a number field `K` at an infinite + - `NumberField.InfinitePlace.Completion` : the completion of a number field `K` at an infinite place, obtained by completing `K` with respect to the absolute value associated to the infinite place. - `NumberField.InfinitePlace.Completion.extensionEmbedding` : the embedding `v.embedding : K →+* ℂ` - extended to `v.completion →+* ℂ`. + extended to `v.Completion →+* ℂ`. - `NumberField.InfinitePlace.Completion.extensionEmbedding_of_isReal` : if the infinite place `v` is real, then this extends the embedding `v.embedding_of_isReal : K →+* ℝ` to - `v.completion →+* ℝ`. + `v.Completion →+* ℝ`. - `NumberField.InfinitePlace.Completion.equiv_real_of_isReal` : the ring isomorphism - `v.completion ≃+* ℝ` when `v` is a real infinite place; the forward direction of this is + `v.Completion ≃+* ℝ` when `v` is a real infinite place; the forward direction of this is `extensionEmbedding_of_isReal`. - `NumberField.InfinitePlace.Completion.equiv_complex_of_isComplex` : the ring isomorphism - `v.completion ≃+* ℂ` when `v` is a complex infinite place; the forward direction of this is + `v.Completion ≃+* ℂ` when `v` is a complex infinite place; the forward direction of this is `extensionEmbedding`. ## Main results - `NumberField.Completion.locallyCompactSpace` : the completion of a number field at an infinite place is locally compact. - - `NumberField.Completion.isometry_extensionEmbedding` : the embedding `v.completion →+* ℂ` is + - `NumberField.Completion.isometry_extensionEmbedding` : the embedding `v.Completion →+* ℂ` is an isometry. See also `isometry_extensionEmbedding_of_isReal` for the corresponding result on - `v.completion →+* ℝ` when `v` is real. + `v.Completion →+* ℝ` when `v` is real. - `NumberField.Completion.bijective_extensionEmbedding_of_isComplex` : the embedding - `v.completion →+* ℂ` is bijective when `v` is complex. See also - `bijective_extensionEmebdding_of_isReal` for the corresponding result for `v.completion →+* ℝ` + `v.Completion →+* ℂ` is bijective when `v` is complex. See also + `bijective_extensionEmebdding_of_isReal` for the corresponding result for `v.Completion →+* ℝ` when `v` is real. ## Tags @@ -62,42 +62,44 @@ open AbsoluteValue.Completion variable {K : Type*} [Field K] (v : InfinitePlace K) /-- The completion of a number field at an infinite place. -/ -abbrev completion := v.1.completion +abbrev Completion := v.1.Completion + +@[deprecated (since := "2024-12-01")] alias completion := Completion namespace Completion -instance : NormedField v.completion := +instance : NormedField v.Completion := letI := (WithAbs.isUniformInducing_of_comp v.norm_embedding_eq).completableTopField UniformSpace.Completion.instNormedFieldOfCompletableTopField (WithAbs v.1) lemma norm_coe (x : WithAbs v.1) : - ‖(x : v.completion)‖ = v (WithAbs.equiv v.1 x) := + ‖(x : v.Completion)‖ = v (WithAbs.equiv v.1 x) := UniformSpace.Completion.norm_coe x -instance : Algebra K v.completion := - inferInstanceAs <| Algebra (WithAbs v.1) v.1.completion +instance : Algebra K v.Completion := + inferInstanceAs <| Algebra (WithAbs v.1) v.1.Completion /-- The coercion from the rationals to its completion along an infinite place is `Rat.cast`. -/ lemma WithAbs.ratCast_equiv (v : InfinitePlace ℚ) (x : WithAbs v.1) : - Rat.cast (WithAbs.equiv _ x) = (x : v.completion) := + Rat.cast (WithAbs.equiv _ x) = (x : v.Completion) := (eq_ratCast (UniformSpace.Completion.coeRingHom.comp (WithAbs.ringEquiv v.1).symm.toRingHom) x).symm lemma Rat.norm_infinitePlace_completion (v : InfinitePlace ℚ) (x : ℚ) : - ‖(x : v.completion)‖ = |x| := by + ‖(x : v.Completion)‖ = |x| := by rw [← (WithAbs.equiv v.1).apply_symm_apply x, WithAbs.ratCast_equiv, norm_coe, (WithAbs.equiv v.1).apply_symm_apply, Rat.infinitePlace_apply] /-- The completion of a number field at an infinite place is locally compact. -/ -instance locallyCompactSpace : LocallyCompactSpace (v.completion) := +instance locallyCompactSpace : LocallyCompactSpace (v.Completion) := AbsoluteValue.Completion.locallyCompactSpace v.norm_embedding_eq -/-- The embedding associated to an infinite place extended to an embedding `v.completion →+* ℂ`. -/ -def extensionEmbedding : v.completion →+* ℂ := extensionEmbedding_of_comp v.norm_embedding_eq +/-- The embedding associated to an infinite place extended to an embedding `v.Completion →+* ℂ`. -/ +def extensionEmbedding : v.Completion →+* ℂ := extensionEmbedding_of_comp v.norm_embedding_eq -/-- The embedding `K →+* ℝ` associated to a real infinite place extended to `v.completion →+* ℝ`. -/ -def extensionEmbedding_of_isReal {v : InfinitePlace K} (hv : IsReal v) : v.completion →+* ℝ := +/-- The embedding `K →+* ℝ` associated to a real infinite place extended to `v.Completion →+* ℝ`. -/ +def extensionEmbedding_of_isReal {v : InfinitePlace K} (hv : IsReal v) : v.Completion →+* ℝ := extensionEmbedding_of_comp <| v.norm_embedding_of_isReal hv @[simp] @@ -109,20 +111,20 @@ theorem extensionEmbedding_of_isReal_coe {v : InfinitePlace K} (hv : IsReal v) ( extensionEmbedding_of_isReal hv x = embedding_of_isReal hv x := extensionEmbedding_of_comp_coe (v.norm_embedding_of_isReal hv) x -/-- The embedding `v.completion →+* ℂ` is an isometry. -/ +/-- The embedding `v.Completion →+* ℂ` is an isometry. -/ theorem isometry_extensionEmbedding : Isometry (extensionEmbedding v) := Isometry.of_dist_eq (extensionEmbedding_dist_eq_of_comp v.norm_embedding_eq) -/-- The embedding `v.completion →+* ℝ` at a real infinite palce is an isometry. -/ +/-- The embedding `v.Completion →+* ℝ` at a real infinite palce is an isometry. -/ theorem isometry_extensionEmbedding_of_isReal {v : InfinitePlace K} (hv : IsReal v) : Isometry (extensionEmbedding_of_isReal hv) := Isometry.of_dist_eq (extensionEmbedding_dist_eq_of_comp <| v.norm_embedding_of_isReal hv) -/-- The embedding `v.completion →+* ℂ` has closed image inside `ℂ`. -/ +/-- The embedding `v.Completion →+* ℂ` has closed image inside `ℂ`. -/ theorem isClosed_image_extensionEmbedding : IsClosed (Set.range (extensionEmbedding v)) := (isClosedEmbedding_extensionEmbedding_of_comp v.norm_embedding_eq).isClosed_range -/-- The embedding `v.completion →+* ℝ` associated to a real infinite place has closed image +/-- The embedding `v.Completion →+* ℝ` associated to a real infinite place has closed image inside `ℝ`. -/ theorem isClosed_image_extensionEmbedding_of_isReal {v : InfinitePlace K} (hv : IsReal v) : IsClosed (Set.range (extensionEmbedding_of_isReal hv)) := @@ -136,48 +138,48 @@ theorem subfield_ne_real_of_isComplex {v : InfinitePlace K} (hv : IsComplex v) : obtain ⟨r, hr⟩ := hv ▸ extensionEmbedding_coe v x ▸ RingHom.mem_fieldRange_self _ _ simp only [ComplexEmbedding.conjugate_coe_eq, ← hr, Complex.ofRealHom_eq_coe, Complex.conj_ofReal] -/-- If `v` is a complex infinite place, then the embedding `v.completion →+* ℂ` is surjective. -/ +/-- If `v` is a complex infinite place, then the embedding `v.Completion →+* ℂ` is surjective. -/ theorem surjective_extensionEmbedding_of_isComplex {v : InfinitePlace K} (hv : IsComplex v) : Function.Surjective (extensionEmbedding v) := by rw [← RingHom.fieldRange_eq_top_iff] exact (Complex.subfield_eq_of_closed <| isClosed_image_extensionEmbedding v).resolve_left <| subfield_ne_real_of_isComplex hv -/-- If `v` is a complex infinite place, then the embedding `v.completion →+* ℂ` is bijective. -/ +/-- If `v` is a complex infinite place, then the embedding `v.Completion →+* ℂ` is bijective. -/ theorem bijective_extensionEmbedding_of_isComplex {v : InfinitePlace K} (hv : IsComplex v) : Function.Bijective (extensionEmbedding v) := ⟨(extensionEmbedding v).injective, surjective_extensionEmbedding_of_isComplex hv⟩ -/-- The ring isomorphism `v.completion ≃+* ℂ`, when `v` is complex, given by the bijection -`v.completion →+* ℂ`. -/ +/-- The ring isomorphism `v.Completion ≃+* ℂ`, when `v` is complex, given by the bijection +`v.Completion →+* ℂ`. -/ def ringEquiv_complex_of_isComplex {v : InfinitePlace K} (hv : IsComplex v) : - v.completion ≃+* ℂ := + v.Completion ≃+* ℂ := RingEquiv.ofBijective _ (bijective_extensionEmbedding_of_isComplex hv) -/-- If the infinite place `v` is complex, then `v.completion` is isometric to `ℂ`. -/ +/-- If the infinite place `v` is complex, then `v.Completion` is isometric to `ℂ`. -/ def isometryEquiv_complex_of_isComplex {v : InfinitePlace K} (hv : IsComplex v) : - v.completion ≃ᵢ ℂ where + v.Completion ≃ᵢ ℂ where toEquiv := ringEquiv_complex_of_isComplex hv isometry_toFun := isometry_extensionEmbedding v -/-- If `v` is a real infinite place, then the embedding `v.completion →+* ℝ` is surjective. -/ +/-- If `v` is a real infinite place, then the embedding `v.Completion →+* ℝ` is surjective. -/ theorem surjective_extensionEmbedding_of_isReal {v : InfinitePlace K} (hv : IsReal v) : Function.Surjective (extensionEmbedding_of_isReal hv) := by rw [← RingHom.fieldRange_eq_top_iff, ← Real.subfield_eq_of_closed] exact isClosed_image_extensionEmbedding_of_isReal hv -/-- If `v` is a real infinite place, then the embedding `v.completion →+* ℝ` is bijective. -/ +/-- If `v` is a real infinite place, then the embedding `v.Completion →+* ℝ` is bijective. -/ theorem bijective_extensionEmbedding_of_isReal {v : InfinitePlace K} (hv : IsReal v) : Function.Bijective (extensionEmbedding_of_isReal hv) := ⟨(extensionEmbedding_of_isReal hv).injective, surjective_extensionEmbedding_of_isReal hv⟩ -/-- The ring isomorphism `v.completion ≃+* ℝ`, when `v` is real, given by the bijection -`v.completion →+* ℝ`. -/ -def ringEquiv_real_of_isReal {v : InfinitePlace K} (hv : IsReal v) : v.completion ≃+* ℝ := +/-- The ring isomorphism `v.Completion ≃+* ℝ`, when `v` is real, given by the bijection +`v.Completion →+* ℝ`. -/ +def ringEquiv_real_of_isReal {v : InfinitePlace K} (hv : IsReal v) : v.Completion ≃+* ℝ := RingEquiv.ofBijective _ (bijective_extensionEmbedding_of_isReal hv) -/-- If the infinite place `v` is real, then `v.completion` is isometric to `ℝ`. -/ -def isometryEquiv_real_of_isReal {v : InfinitePlace K} (hv : IsReal v) : v.completion ≃ᵢ ℝ where +/-- If the infinite place `v` is real, then `v.Completion` is isometric to `ℝ`. -/ +def isometryEquiv_real_of_isReal {v : InfinitePlace K} (hv : IsReal v) : v.Completion ≃ᵢ ℝ where toEquiv := ringEquiv_real_of_isReal hv isometry_toFun := isometry_extensionEmbedding_of_isReal hv From 653f17d0c439d2d9d55797995521660adc6df994 Mon Sep 17 00:00:00 2001 From: su00000 Date: Sat, 7 Dec 2024 06:05:49 +0000 Subject: [PATCH 561/829] feat(RingTheory/LocalProperties): flatness is a local property (#19085) This PR proves that flat is a local property. Main results: 1. `flat_of_localization_maximal` : Let `M` be a module over `CommRing R`. If the localization of `M` at each maximal ideal `P` is flat over `R_p`, then `M` is flat over `R`. 2. `flat_of_localization_span` : Let `M` be a module over `CommRing R` and `S` be a set that spans `R`. If the localization of `M` at `s : S` is flat over `Localization.Away s`, then `M` is flat over `R`. 3. `flat_iff_of_localization` : Let `R_p` a localization of `CommRing R` and `M` be a module over `R_p`. Then `M` is flat over `R` if and only if `M` is flat over `R_p` - [x] depends on: #19080 Co-authored-by: Junyan Xu --- .../Algebra/Module/LocalizedModule/Basic.lean | 5 ++ Mathlib/RingTheory/Flat/Localization.lean | 82 +++++++++++++++++-- 2 files changed, 82 insertions(+), 5 deletions(-) diff --git a/Mathlib/Algebra/Module/LocalizedModule/Basic.lean b/Mathlib/Algebra/Module/LocalizedModule/Basic.lean index efe62ee72caee..b7f0ab08698d3 100644 --- a/Mathlib/Algebra/Module/LocalizedModule/Basic.lean +++ b/Mathlib/Algebra/Module/LocalizedModule/Basic.lean @@ -914,6 +914,11 @@ theorem is_universal : ∃! l : M' →ₗ[R] M'', l.comp f = g := fun g h => ⟨lift S f g h, lift_comp S f g h, fun l hl => (lift_unique S f g h l hl).symm⟩ +theorem linearMap_ext {N N'} [AddCommMonoid N] [Module R N] [AddCommMonoid N'] [Module R N'] + (f' : N →ₗ[R] N') [IsLocalizedModule S f'] ⦃g g' : M' →ₗ[R] N'⦄ + (h : g ∘ₗ f = g' ∘ₗ f) : g = g' := + (is_universal S f _ <| map_units f').unique h rfl + theorem ringHom_ext (map_unit : ∀ x : S, IsUnit ((algebraMap R (Module.End R M'')) x)) ⦃j k : M' →ₗ[R] M''⦄ (h : j.comp f = k.comp f) : j = k := by rw [← lift_unique S f (k.comp f) map_unit j h, lift_unique] diff --git a/Mathlib/RingTheory/Flat/Localization.lean b/Mathlib/RingTheory/Flat/Localization.lean index 9269fb9b8bd69..2de407505a926 100644 --- a/Mathlib/RingTheory/Flat/Localization.lean +++ b/Mathlib/RingTheory/Flat/Localization.lean @@ -3,21 +3,30 @@ Copyright (c) 2024 Junyan Xu. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Junyan Xu -/ -import Mathlib.Algebra.Module.LocalizedModule.Submodule -import Mathlib.RingTheory.Flat.Basic -import Mathlib.RingTheory.Localization.BaseChange +import Mathlib.RingTheory.Flat.Stability +import Mathlib.RingTheory.LocalProperties.Exactness /-! # Flatness and localization -In this file we show that localizations are flat, and flatness is a local property (TODO). +In this file we show that localizations are flat, and flatness is a local property. ## Main result -- `IsLocalization.flat`: a localization of a commutative ring is flat over it. +* `IsLocalization.flat`: a localization of a commutative ring is flat over it. +* `Module.flat_iff_of_isLocalization` : Let `Rₚ` a localization of a commutative ring `R` + and `M` be a module over `Rₚ`. Then `M` is flat over `R` if and only if `M` is flat over `Rₚ`. +* `Module.flat_of_isLocalized_maximal` : Let `M` be a module over a commutative ring `R`. + If the localization of `M` at each maximal ideal `P` is flat over `Rₚ`, then `M` is flat over `R`. +* `Module.flat_of_isLocalized_span` : Let `M` be a module over a commutative ring `R` + and `S` be a set that spans `R`. If the localization of `M` at each `s : S` is flat + over `Localization.Away s`, then `M` is flat over `R`. -/ +open IsLocalizedModule LocalizedModule LinearMap TensorProduct + variable {R : Type*} (S : Type*) [CommRing R] [CommRing S] [Algebra R S] variable (p : Submonoid R) [IsLocalization p S] +variable (M : Type*) [AddCommGroup M] [Module R M] [Module S M] [IsScalarTower R S M] include p in theorem IsLocalization.flat : Module.Flat R S := @@ -30,3 +39,66 @@ theorem IsLocalization.flat : Module.Flat R S := simpa [this, - Subtype.val_injective] using Subtype.val_injective instance Localization.flat : Module.Flat R (Localization p) := IsLocalization.flat _ p + +namespace Module + +include p in +theorem flat_iff_of_isLocalization : Flat S M ↔ Flat R M := + have := isLocalizedModule_id p M S + have := IsLocalization.flat S p + ⟨fun _ ↦ .trans R S M, fun _ ↦ .of_isLocalizedModule S p .id⟩ + +private lemma aux (I : Ideal R) (s : Submonoid R) : + have hM := isBaseChange s (Localization s) (mkLinearMap s M) + have hIM := isBaseChange s (Localization s) (mkLinearMap s (I ⊗[R] M)) + let e := (hIM.equiv.restrictScalars R).symm ≪≫ₗ + leftComm _ _ _ _ ≪≫ₗ (hM.equiv.restrictScalars R).lTensor I + LocalizedModule.map s (TensorProduct.lift <| lsmul R M ∘ₗ I.subtype) = + TensorProduct.lift (lsmul R (LocalizedModule s M) ∘ₗ I.subtype) ∘ₗ e.toLinearMap := by + refine linearMap_ext s (mkLinearMap s _) (mkLinearMap s _) ?_ + ext m i + rw [AlgebraTensorModule.curry_apply, curry_apply, restrictScalars_apply, LinearMap.comp_apply, + restrictScalars_apply, mkLinearMap_apply, AlgebraTensorModule.curry_apply, curry_apply, + restrictScalars_apply] + simpa [-mkLinearMap_apply, IsBaseChange.equiv_symm_apply, IsBaseChange.equiv_tmul] using + (mkLinearMap_apply _ _ _).symm.trans (map_smul (mkLinearMap s M) _ _) + +variable (Mₚ : ∀ (P : Ideal R) [P.IsMaximal], Type*) + [∀ (P : Ideal R) [P.IsMaximal], AddCommGroup (Mₚ P)] + [∀ (P : Ideal R) [P.IsMaximal], Module R (Mₚ P)] + (f : ∀ (P : Ideal R) [P.IsMaximal], M →ₗ[R] Mₚ P) + [∀ (P : Ideal R) [P.IsMaximal], IsLocalizedModule P.primeCompl (f P)] + +theorem flat_of_localized_maximal + (h : ∀ (J : Ideal R) [J.IsMaximal], Flat R (LocalizedModule J.primeCompl M)) : + Flat R M := + (flat_iff _ _).mpr fun I fg ↦ injective_of_localized_maximal _ fun J hJ ↦ by + rw [← LinearMap.coe_restrictScalars R, aux] + simpa using (flat_iff _ _).mp (h J) fg + +include f in +theorem flat_of_isLocalized_maixmal (H : ∀ (P : Ideal R) [P.IsMaximal], Flat R (Mₚ P)) : + Module.Flat R M := + flat_of_localized_maximal M fun P _ ↦ .of_linearEquiv _ _ _ (iso _ (f P)) + +variable (s : Set R) (spn : Ideal.span s = ⊤) + (Mₛ : ∀ _ : s, Type*) + [∀ r : s, AddCommGroup (Mₛ r)] + [∀ r : s, Module R (Mₛ r)] + (g : ∀ r : s, M →ₗ[R] Mₛ r) + [∀ r : s, IsLocalizedModule (.powers r.1) (g r)] +include spn + +theorem flat_of_localized_span + (h : ∀ r : s, Flat R (LocalizedModule (.powers r.1) M)) : + Flat R M := + (Module.flat_iff _ _).mpr fun I fg ↦ injective_of_localized_span s spn _ fun r ↦ by + rw [← LinearMap.coe_restrictScalars R, aux] + simpa using (Module.flat_iff _ _).mp (h r) fg + +include g in +theorem flat_of_isLocalized_span (H : ∀ r : s, Module.Flat R (Mₛ r)) : + Module.Flat R M := + flat_of_localized_span M s spn fun r ↦ .of_linearEquiv _ _ _ (iso _ (g r)) + +end Module From 37305b9e5eff20c7b426a9905ccd04b387203bf1 Mon Sep 17 00:00:00 2001 From: grunweg Date: Sat, 7 Dec 2024 09:19:58 +0000 Subject: [PATCH 562/829] perf: squeeze some really slow `simp`s (#19743) See the individual commit messages for details: each simp was particular slow, and is much faster now. --- Mathlib/Algebra/GroupWithZero/Pointwise/Finset.lean | 10 ++++++++-- Mathlib/Algebra/Lie/Weights/Chain.lean | 7 +++++-- Mathlib/FieldTheory/AxGrothendieck.lean | 6 +++++- Mathlib/RingTheory/Kaehler/Basic.lean | 5 ++++- 4 files changed, 22 insertions(+), 6 deletions(-) diff --git a/Mathlib/Algebra/GroupWithZero/Pointwise/Finset.lean b/Mathlib/Algebra/GroupWithZero/Pointwise/Finset.lean index c112e0a96c541..0de5a19d7fbc0 100644 --- a/Mathlib/Algebra/GroupWithZero/Pointwise/Finset.lean +++ b/Mathlib/Algebra/GroupWithZero/Pointwise/Finset.lean @@ -185,12 +185,18 @@ open scoped RightActions @[simp] lemma inv_smul_finset_distrib₀ (a : α) (s : Finset α) : (a • s)⁻¹ = s⁻¹ <• a⁻¹ := by obtain rfl | ha := eq_or_ne a 0 · obtain rfl | hs := s.eq_empty_or_nonempty <;> simp [*] - · ext; simp [← inv_smul_mem_iff₀, *] + -- was `simp` and very slow (https://github.com/leanprover-community/mathlib4/issues/19751) + · ext; simp only [mem_inv', ne_eq, not_false_eq_true, ← inv_smul_mem_iff₀, smul_eq_mul, + MulOpposite.op_inv, inv_eq_zero, MulOpposite.op_eq_zero_iff, inv_inv, + MulOpposite.smul_eq_mul_unop, MulOpposite.unop_op, mul_inv_rev, ha] @[simp] lemma inv_op_smul_finset_distrib₀ (a : α) (s : Finset α) : (s <• a)⁻¹ = a⁻¹ • s⁻¹ := by obtain rfl | ha := eq_or_ne a 0 · obtain rfl | hs := s.eq_empty_or_nonempty <;> simp [*] - · ext; simp [← inv_smul_mem_iff₀, *] + -- was `simp` and very slow (https://github.com/leanprover-community/mathlib4/issues/19751) + · ext; simp only [mem_inv', ne_eq, MulOpposite.op_eq_zero_iff, not_false_eq_true, ← + inv_smul_mem_iff₀, MulOpposite.smul_eq_mul_unop, MulOpposite.unop_inv, MulOpposite.unop_op, + inv_eq_zero, inv_inv, smul_eq_mul, mul_inv_rev, ha] end GroupWithZero diff --git a/Mathlib/Algebra/Lie/Weights/Chain.lean b/Mathlib/Algebra/Lie/Weights/Chain.lean index 1acfd102224c7..01285cf51e542 100644 --- a/Mathlib/Algebra/Lie/Weights/Chain.lean +++ b/Mathlib/Algebra/Lie/Weights/Chain.lean @@ -132,8 +132,11 @@ lemma lie_mem_genWeightSpaceChain_of_genWeightSpace_eq_bot_right [LieAlgebra.IsN obtain ⟨k, hk⟩ := k suffices genWeightSpace M ((k + 1) • α + χ) ≤ genWeightSpaceChain M α χ p q by apply this - simpa using (rootSpaceWeightSpaceProduct R L H M α (k • α + χ) ((k + 1) • α + χ) - (by rw [add_smul]; abel) (⟨x, hx⟩ ⊗ₜ ⟨z, hz⟩)).property + -- was `simpa using [...]` and very slow + -- (https://github.com/leanprover-community/mathlib4/issues/19751) + simpa only [zsmul_eq_mul, Int.cast_add, Pi.intCast_def, Int.cast_one] using + (rootSpaceWeightSpaceProduct R L H M α (k • α + χ) ((k + 1) • α + χ) + (by rw [add_smul]; abel) (⟨x, hx⟩ ⊗ₜ ⟨z, hz⟩)).property rw [genWeightSpaceChain] rcases eq_or_ne (k + 1) q with rfl | hk'; · simp only [hq, bot_le] replace hk' : k + 1 ∈ Ioo p q := ⟨by linarith [hk.1], lt_of_le_of_ne hk.2 hk'⟩ diff --git a/Mathlib/FieldTheory/AxGrothendieck.lean b/Mathlib/FieldTheory/AxGrothendieck.lean index 5cfd3ecd0a861..a7bc963365950 100644 --- a/Mathlib/FieldTheory/AxGrothendieck.lean +++ b/Mathlib/FieldTheory/AxGrothendieck.lean @@ -160,7 +160,11 @@ theorem realize_genericPolyMapSurjOnOfInjOn Set.MapsTo, Set.mem_def, injOnAlt, funext_iff, Set.SurjOn, Set.image, setOf, Set.subset_def, Equiv.forall_congr_left (mvPolynomialSupportLEEquiv mons)] simp +singlePass only [← Sum.elim_comp_inl_inr] - simp [Set.mem_def, Function.comp_def] + -- was `simp` and very slow (https://github.com/leanprover-community/mathlib4/issues/19751) + simp only [Function.comp_def, Sum.elim_inl, Sum.elim_inr, Fin.castAdd_zero, Fin.cast_eq_self, + Nat.add_zero, Term.realize_var, Term.realize_relabel, realize_termOfFreeCommRing, + lift_genericPolyMap, Nat.reduceAdd, Fin.isValue, Function.uncurry_apply_pair, Fin.cons_zero, + Fin.cons_one, ↓reduceIte, one_ne_zero] theorem ACF_models_genericPolyMapSurjOnOfInjOn_of_prime [Fintype ι] {p : ℕ} (hp : p.Prime) (φ : ring.Formula (α ⊕ ι)) (mons : ι → Finset (ι →₀ ℕ)) : diff --git a/Mathlib/RingTheory/Kaehler/Basic.lean b/Mathlib/RingTheory/Kaehler/Basic.lean index 4fa13a26c251e..d668bcecc861a 100644 --- a/Mathlib/RingTheory/Kaehler/Basic.lean +++ b/Mathlib/RingTheory/Kaehler/Basic.lean @@ -839,7 +839,8 @@ theorem KaehlerDifferential.range_kerCotangentToTensor constructor · rintro ⟨x, rfl⟩ obtain ⟨x, rfl⟩ := Ideal.toCotangent_surjective _ x - simp [kerCotangentToTensor_toCotangent, RingHom.mem_ker.mp x.2] + simp only [kerCotangentToTensor_toCotangent, Submodule.restrictScalars_mem, LinearMap.mem_ker, + mapBaseChange_tmul, map_D, RingHom.mem_ker.mp x.2, map_zero, smul_zero] · intro hx obtain ⟨x, rfl⟩ := LinearMap.rTensor_surjective (Ω[A⁄R]) (g := Algebra.linearMap A B) h x obtain ⟨x, rfl⟩ := (TensorProduct.lid _ _).symm.surjective x @@ -869,6 +870,8 @@ theorem KaehlerDifferential.range_kerCotangentToTensor simp [RingHom.mem_ker, ha, this.2] · simp only [map_sum, LinearMapClass.map_smul, kerCotangentToTensor_toCotangent, map_sub] simp_rw [← TensorProduct.tmul_smul] + -- was `simp [kerCotangentToTensor_toCotangent, RingHom.mem_ker.mp x.2]` and very slow + -- (https://github.com/leanprover-community/mathlib4/issues/19751) simp only [smul_sub, TensorProduct.tmul_sub, Finset.sum_sub_distrib, ← TensorProduct.tmul_sum, ← Finset.sum_smul, Finset.sum_attach, sub_eq_self, Finset.sum_attach (f := fun i ↦ x i • KaehlerDifferential.D R A i)] From 21d992dd68e83a90bc1b22b677f90246efaf214a Mon Sep 17 00:00:00 2001 From: Fabrizio Barroero Date: Sat, 7 Dec 2024 12:21:14 +0000 Subject: [PATCH 563/829] feat(NumberTheory/NumberField/FinitePlaces): the finite places of a number field (#19667) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This file defines finite places of a number field `K` as absolute values coming from an embedding into a completion of `K` associated to a non-zero prime ideal of `𝓞 K` with some basic APIs. This is part of a project which aims to define heights. The next step is the product formula. --- Mathlib.lean | 1 + .../NumberField/FinitePlaces.lean | 236 ++++++++++++++++++ 2 files changed, 237 insertions(+) create mode 100644 Mathlib/NumberTheory/NumberField/FinitePlaces.lean diff --git a/Mathlib.lean b/Mathlib.lean index 9269de2ed24fa..e1579076cc334 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3883,6 +3883,7 @@ import Mathlib.NumberTheory.NumberField.Discriminant.Basic import Mathlib.NumberTheory.NumberField.Discriminant.Defs import Mathlib.NumberTheory.NumberField.Embeddings import Mathlib.NumberTheory.NumberField.EquivReindex +import Mathlib.NumberTheory.NumberField.FinitePlaces import Mathlib.NumberTheory.NumberField.FractionalIdeal import Mathlib.NumberTheory.NumberField.House import Mathlib.NumberTheory.NumberField.Norm diff --git a/Mathlib/NumberTheory/NumberField/FinitePlaces.lean b/Mathlib/NumberTheory/NumberField/FinitePlaces.lean new file mode 100644 index 0000000000000..834648c1e128e --- /dev/null +++ b/Mathlib/NumberTheory/NumberField/FinitePlaces.lean @@ -0,0 +1,236 @@ +/- +Copyright (c) 2024 Fabrizio Barroero. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Fabrizio Barroero +-/ +import Mathlib.Data.Int.WithZero +import Mathlib.NumberTheory.NumberField.Embeddings +import Mathlib.RingTheory.DedekindDomain.AdicValuation +import Mathlib.RingTheory.DedekindDomain.Factorization +import Mathlib.RingTheory.Ideal.Norm.AbsNorm +import Mathlib.Topology.Algebra.Valued.NormedValued + +/-! +# Finite places of number fields +This file defines finite places of a number field `K` as absolute values coming from an embedding +into a completion of `K` associated to a non-zero prime ideal of `𝓞 K`. + +## Main Definitions and Results +* `NumberField.vadicAbv`: a `v`-adic absolute value on `K`. +* `NumberField.FinitePlace`: the type of finite places of a number field `K`. +* `NumberField.FinitePlace.mulSupport_finite`: the `v`-adic absolute value of a non-zero element of +`K` is different from 1 for at most finitely many `v`. + +## Tags +number field, places, finite places +-/ + +open Ideal IsDedekindDomain HeightOneSpectrum NumberField WithZeroMulInt + +namespace NumberField + +section absoluteValue + +variable {K : Type*} [Field K] [NumberField K] (v : HeightOneSpectrum (𝓞 K)) + +/-- The norm of a maximal ideal as an element of `ℝ≥0` is `> 1` -/ +lemma one_lt_norm : 1 < (absNorm v.asIdeal : NNReal) := by + norm_cast + by_contra! h + apply IsPrime.ne_top v.isPrime + rw [← absNorm_eq_one_iff] + have : 0 < absNorm v.asIdeal := by + rw [Nat.pos_iff_ne_zero, absNorm_ne_zero_iff] + exact (v.asIdeal.fintypeQuotientOfFreeOfNeBot v.ne_bot).finite + omega + +private lemma norm_ne_zero : (absNorm v.asIdeal : NNReal) ≠ 0 := ne_zero_of_lt (one_lt_norm v) + +/-- The `v`-adic absolute value on `K` defined as the norm of `v` raised to negative `v`-adic +valuation.-/ +noncomputable def vadicAbv : AbsoluteValue K ℝ where + toFun x := toNNReal (norm_ne_zero v) (v.valuation x) + map_mul' _ _ := by simp only [_root_.map_mul, NNReal.coe_mul] + nonneg' _ := NNReal.zero_le_coe + eq_zero' _ := by simp only [NNReal.coe_eq_zero, map_eq_zero] + add_le' x y := by + -- the triangle inequality is implied by the ultrametric one + apply le_trans _ <| max_le_add_of_nonneg (zero_le ((toNNReal (norm_ne_zero v)) (v.valuation x))) + (zero_le ((toNNReal (norm_ne_zero v)) (v.valuation y))) + have h_mono := (toNNReal_strictMono (one_lt_norm v)).monotone + rw [← h_mono.map_max] --max goes inside withZeroMultIntToNNReal + exact h_mono (v.valuation.map_add x y) + +theorem vadicAbv_def {x : K} : vadicAbv v x = toNNReal (norm_ne_zero v) (v.valuation x) := rfl + +end absoluteValue + +section FinitePlace +variable {K : Type*} [Field K] [NumberField K] (v : HeightOneSpectrum (𝓞 K)) + +/-- The embedding of a number field inside its completion with respect to `v`. -/ +def embedding : K →+* adicCompletion K v := + @UniformSpace.Completion.coeRingHom K _ v.adicValued.toUniformSpace _ _ + +noncomputable instance instRankOneValuedAdicCompletion : + Valuation.RankOne (valuedAdicCompletion K v).v where + hom := { + toFun := toNNReal (norm_ne_zero v) + map_zero' := rfl + map_one' := rfl + map_mul' := MonoidWithZeroHom.map_mul (toNNReal (norm_ne_zero v)) + } + strictMono' := toNNReal_strictMono (one_lt_norm v) + nontrivial' := by + rcases Submodule.exists_mem_ne_zero_of_ne_bot v.ne_bot with ⟨x, hx1, hx2⟩ + use (x : K) + rw [valuedAdicCompletion_eq_valuation' v (x : K)] + constructor + · simpa only [ne_eq, map_eq_zero, NoZeroSMulDivisors.algebraMap_eq_zero_iff] + · apply ne_of_lt + rw [valuation_eq_intValuationDef, intValuation_lt_one_iff_dvd] + exact dvd_span_singleton.mpr hx1 + +/-- The `v`-adic completion of `K` is a normed field. -/ +noncomputable instance instNormedFieldValuedAdicCompletion : NormedField (adicCompletion K v) := + Valued.toNormedField (adicCompletion K v) (WithZero (Multiplicative ℤ)) + +/-- A finite place of a number field `K` is a place associated to an embedding into a completion +with respect to a maximal ideal. -/ +def FinitePlace (K : Type*) [Field K] [NumberField K] := + {w : AbsoluteValue K ℝ // ∃ v : HeightOneSpectrum (𝓞 K), place (embedding v) = w} + +/-- Return the finite place defined by a maximal ideal `v`. -/ +noncomputable def FinitePlace.mk (v : HeightOneSpectrum (𝓞 K)) : FinitePlace K := + ⟨place (embedding v), ⟨v, rfl⟩⟩ + +lemma toNNReal_Valued_eq_vadicAbv (x : K) : + toNNReal (norm_ne_zero v) (Valued.v (self:=v.adicValued) x) = vadicAbv v x := rfl + +/-- The norm of the image after the embedding associated to `v` is equal to the `v`-adic absolute +value. -/ +theorem FinitePlace.norm_def (x : K) : ‖embedding v x‖ = vadicAbv v x := by + simp only [NormedField.toNorm, instNormedFieldValuedAdicCompletion, Valued.toNormedField, + instFieldAdicCompletion, Valued.norm, Valuation.RankOne.hom, MonoidWithZeroHom.coe_mk, + ZeroHom.coe_mk, embedding, UniformSpace.Completion.coeRingHom, RingHom.coe_mk, MonoidHom.coe_mk, + OneHom.coe_mk, Valued.valuedCompletion_apply, toNNReal_Valued_eq_vadicAbv] + +/-- The norm of the image after the embedding associated to `v` is equal to the norm of `v` raised +to the power of the `v`-adic valuation. -/ +theorem FinitePlace.norm_def' (x : K) : ‖embedding v x‖ = toNNReal (norm_ne_zero v) + (v.valuation x) := by + rw [norm_def, vadicAbv_def] + +/-- The norm of the image after the embedding associated to `v` is equal to the norm of `v` raised +to the power of the `v`-adic valuation for integers. -/ +theorem FinitePlace.norm_def_int (x : 𝓞 K) : ‖embedding v x‖ = toNNReal (norm_ne_zero v) + (v.intValuationDef x) := by + rw [norm_def, vadicAbv_def, valuation_eq_intValuationDef] + +open FinitePlace + +/-- The `v`-adic norm of an integer is at most 1. -/ +theorem norm_le_one (x : 𝓞 K) : ‖embedding v x‖ ≤ 1 := by + rw [norm_def', NNReal.coe_le_one, toNNReal_le_one_iff (one_lt_norm v)] + exact valuation_le_one v x + +/-- The `v`-adic norm of an integer is 1 if and only if it is not in the ideal. -/ +theorem norm_eq_one_iff_not_mem (x : 𝓞 K) : ‖(embedding v) x‖ = 1 ↔ x ∉ v.asIdeal := by + rw [norm_def_int, NNReal.coe_eq_one, toNNReal_eq_one_iff (v.intValuationDef x) + (norm_ne_zero v) (one_lt_norm v).ne', ← dvd_span_singleton, + ← intValuation_lt_one_iff_dvd, not_lt] + exact (intValuation_le_one v x).ge_iff_eq.symm + +/-- The `v`-adic norm of an integer is less than 1 if and only if it is in the ideal. -/ +theorem norm_lt_one_iff_mem (x : 𝓞 K) : ‖embedding v x‖ < 1 ↔ x ∈ v.asIdeal := by + rw [norm_def_int, NNReal.coe_lt_one, toNNReal_lt_one_iff (one_lt_norm v), + intValuation_lt_one_iff_dvd] + exact dvd_span_singleton + +end FinitePlace + +namespace FinitePlace +variable {K : Type*} [Field K] [NumberField K] + +instance : FunLike (FinitePlace K) K ℝ where + coe w x := w.1 x + coe_injective' _ _ h := Subtype.eq (AbsoluteValue.ext <| congr_fun h) + +instance : MonoidWithZeroHomClass (FinitePlace K) K ℝ where + map_mul w := w.1.map_mul + map_one w := w.1.map_one + map_zero w := w.1.map_zero + +instance : NonnegHomClass (FinitePlace K) K ℝ where + apply_nonneg w := w.1.nonneg + +@[simp] +theorem apply (v : HeightOneSpectrum (𝓞 K)) (x : K) : mk v x = ‖embedding v x‖ := rfl + +/-- For a finite place `w`, return a maximal ideal `v` such that `w = finite_place v` . -/ +noncomputable def maximalIdeal (w : FinitePlace K) : HeightOneSpectrum (𝓞 K) := w.2.choose + +@[simp] +theorem mk_maximalIdeal (w : FinitePlace K) : mk (maximalIdeal w) = w := Subtype.ext w.2.choose_spec + +@[simp] +theorem norm_embedding_eq (w : FinitePlace K) (x : K) : + ‖embedding (maximalIdeal w) x‖ = w x := by + conv_rhs => rw [← mk_maximalIdeal w, apply] + +theorem pos_iff {w : FinitePlace K} {x : K} : 0 < w x ↔ x ≠ 0 := AbsoluteValue.pos_iff w.1 + +@[simp] +theorem mk_eq_iff {v₁ v₂ : HeightOneSpectrum (𝓞 K)} : mk v₁ = mk v₂ ↔ v₁ = v₂ := by + refine ⟨?_, fun a ↦ by rw [a]⟩ + contrapose! + intro h + rw [DFunLike.ne_iff] + have ⟨x, hx1, hx2⟩ : ∃ x : 𝓞 K, x ∈ v₁.asIdeal ∧ x ∉ v₂.asIdeal := by + by_contra! H + exact h <| HeightOneSpectrum.ext_iff.mpr <| IsMaximal.eq_of_le (isMaximal v₁) IsPrime.ne_top' H + use x + simp only [apply] + rw [← norm_lt_one_iff_mem ] at hx1 + rw [← norm_eq_one_iff_not_mem] at hx2 + linarith + +theorem maximalIdeal_mk (v : HeightOneSpectrum (𝓞 K)) : maximalIdeal (mk v) = v := by + rw [← mk_eq_iff, mk_maximalIdeal] + +lemma maximalIdeal_injective : (fun w : FinitePlace K ↦ maximalIdeal w).Injective := + Function.HasLeftInverse.injective ⟨mk, mk_maximalIdeal⟩ + +lemma maximalIdeal_inj (w₁ w₂ : FinitePlace K) : maximalIdeal w₁ = maximalIdeal w₂ ↔ w₁ = w₂ := + maximalIdeal_injective.eq_iff + +theorem mulSupport_finite_int {x : 𝓞 K} (h_x_nezero : x ≠ 0) : + (Function.mulSupport fun w : FinitePlace K ↦ w x).Finite := by + have (w : FinitePlace K) : w x ≠ 1 ↔ w x < 1 := + ne_iff_lt_iff_le.mpr <| norm_embedding_eq w x ▸ norm_le_one w.maximalIdeal x + simp_rw [Function.mulSupport, this, ← norm_embedding_eq, norm_lt_one_iff_mem, + ← Ideal.dvd_span_singleton] + have h : {v : HeightOneSpectrum (𝓞 K) | v.asIdeal ∣ span {x}}.Finite := by + apply Ideal.finite_factors + simp only [Submodule.zero_eq_bot, ne_eq, span_singleton_eq_bot, h_x_nezero, not_false_eq_true] + have h_inj : Set.InjOn FinitePlace.maximalIdeal {w | w.maximalIdeal.asIdeal ∣ span {x}} := + Function.Injective.injOn maximalIdeal_injective + refine (h.subset ?_).of_finite_image h_inj + simp only [dvd_span_singleton, Set.image_subset_iff, Set.preimage_setOf_eq, subset_refl] + +theorem mulSupport_finite {x : K} (h_x_nezero : x ≠ 0) : + (Function.mulSupport fun w : FinitePlace K ↦ w x).Finite := by + rcases IsFractionRing.div_surjective (A := 𝓞 K) x with ⟨a, b, hb, rfl⟩ + simp_all only [ne_eq, div_eq_zero_iff, NoZeroSMulDivisors.algebraMap_eq_zero_iff, not_or, + map_div₀] + obtain ⟨ha, hb⟩ := h_x_nezero + simp_rw [← RingOfIntegers.coe_eq_algebraMap] + apply ((mulSupport_finite_int ha).union (mulSupport_finite_int hb)).subset + intro w + simp only [Function.mem_mulSupport, ne_eq, Set.mem_union] + contrapose! + simp +contextual only [ne_eq, one_ne_zero, not_false_eq_true, div_self, implies_true] + +end FinitePlace + +end NumberField From fda98ef198c835a5db26d1e5eed308fec0e6f881 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 7 Dec 2024 12:44:08 +0000 Subject: [PATCH 564/829] =?UTF-8?q?feat(Nat):=20`m=20/=20n=20=3D=200=20?= =?UTF-8?q?=E2=86=94=20n=20=3D=200=20=E2=88=A8=20m=20<=20n`=20(#19505)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Strengthen three iffs by moving their precondition to the RHS. From LeanCamCombi --- Mathlib/Algebra/Order/Antidiag/Pi.lean | 3 +- .../DivisionPolynomial/Degree.lean | 2 +- Mathlib/Data/Nat/Defs.lean | 36 +++++++++---------- Mathlib/Data/Nat/Digits.lean | 7 ++-- Mathlib/Data/Nat/Factorization/Basic.lean | 2 +- Mathlib/Data/Nat/Squarefree.lean | 3 +- Mathlib/Data/ZMod/Basic.lean | 2 +- 7 files changed, 27 insertions(+), 28 deletions(-) diff --git a/Mathlib/Algebra/Order/Antidiag/Pi.lean b/Mathlib/Algebra/Order/Antidiag/Pi.lean index f2c11d7931ca5..aa2f6b9fb2e4a 100644 --- a/Mathlib/Algebra/Order/Antidiag/Pi.lean +++ b/Mathlib/Algebra/Order/Antidiag/Pi.lean @@ -223,7 +223,8 @@ lemma nsmul_piAntidiag [DecidableEq (ι → ℕ)] (s : Finset ι) (m : ℕ) {n : · rw [not_imp_comm.1 (hfsup _) hi] exact dvd_zero _ refine ⟨fun i ↦ f i / n, ?_⟩ - simpa [Nat.sum_div, Nat.div_ne_zero_iff_of_dvd, funext_iff, Nat.mul_div_cancel', ← Nat.sum_div, *] + simp [funext_iff, Nat.mul_div_cancel', ← Nat.sum_div, *] + aesop lemma map_nsmul_piAntidiag (s : Finset ι) (m : ℕ) {n : ℕ} (hn : n ≠ 0) : (piAntidiag s m).map diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Degree.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Degree.lean index ac58cdee1d2ca..c92f4e2898103 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Degree.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Degree.lean @@ -249,7 +249,7 @@ lemma natDegree_preΨ' {n : ℕ} (h : (n : R) ≠ 0) : natDegree_eq_of_le_of_coeff_ne_zero (W.natDegree_preΨ'_le n) <| W.coeff_preΨ'_ne_zero h lemma natDegree_preΨ'_pos {n : ℕ} (hn : 2 < n) (h : (n : R) ≠ 0) : 0 < (W.preΨ' n).natDegree := by - rw [W.natDegree_preΨ' h, Nat.div_pos_iff two_ne_zero] + simp only [W.natDegree_preΨ' h, Nat.div_pos_iff, zero_lt_two, true_and] split_ifs <;> exact Nat.AtLeastTwo.prop.trans <| Nat.sub_le_sub_right (Nat.pow_le_pow_of_le_left hn 2) _ diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index fc518c6dbbc1c..d4a9697c1a97e 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -438,12 +438,23 @@ protected lemma div_le_div_right (h : a ≤ b) : a / c ≤ b / c := lemma lt_of_div_lt_div (h : a / c < b / c) : a < b := Nat.lt_of_not_le fun hab ↦ Nat.not_le_of_lt h <| Nat.div_le_div_right hab -protected lemma div_pos (hba : b ≤ a) (hb : 0 < b) : 0 < a / b := - Nat.pos_of_ne_zero fun h ↦ Nat.lt_irrefl a <| - calc - a = a % b := by simpa [h] using (mod_add_div a b).symm - _ < b := mod_lt a hb - _ ≤ a := hba +@[simp] protected lemma div_eq_zero_iff : a / b = 0 ↔ b = 0 ∨ a < b where + mp h := by + rw [← mod_add_div a b, h, Nat.mul_zero, Nat.add_zero, or_iff_not_imp_left] + exact mod_lt _ ∘ Nat.pos_iff_ne_zero.2 + mpr := by + obtain rfl | hb := eq_or_ne b 0 + · simp + simp only [hb, false_or] + rw [← Nat.mul_right_inj hb, ← Nat.add_left_cancel_iff, mod_add_div] + simp +contextual [mod_eq_of_lt] + +protected lemma div_ne_zero_iff : a / b ≠ 0 ↔ b ≠ 0 ∧ b ≤ a := by simp + +@[simp] protected lemma div_pos_iff : 0 < a / b ↔ 0 < b ∧ b ≤ a := by + simp [Nat.pos_iff_ne_zero] + +protected lemma div_pos (hba : b ≤ a) (hb : 0 < b) : 0 < a / b := Nat.div_pos_iff.2 ⟨hb, hba⟩ lemma lt_mul_of_div_lt (h : a / c < b) (hc : 0 < c) : a < b * c := Nat.lt_of_not_ge <| Nat.not_le_of_gt h ∘ (Nat.le_div_iff_mul_le hc).2 @@ -993,17 +1004,6 @@ lemma div_eq_iff_eq_of_dvd_dvd (hn : n ≠ 0) (ha : a ∣ n) (hb : b ∣ n) : n exact Nat.eq_mul_of_div_eq_right ha h · rw [h] -protected lemma div_eq_zero_iff (hb : 0 < b) : a / b = 0 ↔ a < b where - mp h := by rw [← mod_add_div a b, h, Nat.mul_zero, Nat.add_zero]; exact mod_lt _ hb - mpr h := by rw [← Nat.mul_right_inj (Nat.ne_of_gt hb), ← Nat.add_left_cancel_iff, mod_add_div, - mod_eq_of_lt h, Nat.mul_zero, Nat.add_zero] - -protected lemma div_ne_zero_iff (hb : b ≠ 0) : a / b ≠ 0 ↔ b ≤ a := by - rw [ne_eq, Nat.div_eq_zero_iff (Nat.pos_of_ne_zero hb), not_lt] - -protected lemma div_pos_iff (hb : b ≠ 0) : 0 < a / b ↔ b ≤ a := by - rw [Nat.pos_iff_ne_zero, Nat.div_ne_zero_iff hb] - lemma le_iff_ne_zero_of_dvd (ha : a ≠ 0) (hab : a ∣ b) : a ≤ b ↔ b ≠ 0 where mp := by rw [← Nat.pos_iff_ne_zero] at ha ⊢; exact Nat.lt_of_lt_of_le ha mpr hb := Nat.le_of_dvd (Nat.pos_iff_ne_zero.2 hb) hab @@ -1266,7 +1266,7 @@ lemma dvd_mul_of_div_dvd (h : b ∣ a) (hdiv : a / b ∣ c) : a ∣ b * c := by /-- If a small natural number is divisible by a larger natural number, the small number is zero. -/ lemma eq_zero_of_dvd_of_lt (w : a ∣ b) (h : b < a) : b = 0 := - Nat.eq_zero_of_dvd_of_div_eq_zero w ((Nat.div_eq_zero_iff (lt_of_le_of_lt (zero_le b) h)).mpr h) + Nat.eq_zero_of_dvd_of_div_eq_zero w (by simp [h]) lemma le_of_lt_add_of_dvd (h : a < b + n) : n ∣ a → n ∣ b → a ≤ b := by rintro ⟨a, rfl⟩ ⟨b, rfl⟩ diff --git a/Mathlib/Data/Nat/Digits.lean b/Mathlib/Data/Nat/Digits.lean index f9385aca645c1..5adc574882628 100644 --- a/Mathlib/Data/Nat/Digits.lean +++ b/Mathlib/Data/Nat/Digits.lean @@ -304,8 +304,8 @@ theorem digits_len (b n : ℕ) (hb : 1 < b) (hn : n ≠ 0) : (b.digits n).length induction' n using Nat.strong_induction_on with n IH rw [digits_eq_cons_digits_div hb hn, List.length] by_cases h : n / b = 0 - · have hb0 : b ≠ 0 := (Nat.succ_le_iff.1 hb).ne_bot - simp [h, log_eq_zero_iff, ← Nat.div_eq_zero_iff hb0.bot_lt] + · simp [IH, h] + aesop · have : n / b < n := div_lt_self (Nat.pos_of_ne_zero hn) hb rw [IH _ this h, log_div_base, tsub_add_cancel_of_le] refine Nat.succ_le_of_lt (log_pos hb ?_) @@ -821,8 +821,7 @@ theorem digits_one (b n) (n0 : 0 < n) (nb : n < b) : Nat.digits b n = [n] ∧ 1 have b2 : 1 < b := lt_iff_add_one_le.mpr (le_trans (add_le_add_right (lt_iff_add_one_le.mp n0) 1) nb) refine ⟨?_, b2, n0⟩ - rw [Nat.digits_def' b2 n0, Nat.mod_eq_of_lt nb, - (Nat.div_eq_zero_iff ((zero_le n).trans_lt nb)).2 nb, Nat.digits_zero] + rw [Nat.digits_def' b2 n0, Nat.mod_eq_of_lt nb, Nat.div_eq_zero_iff.2 <| .inr nb, Nat.digits_zero] /- Porting note: this part of the file is tactic related. diff --git a/Mathlib/Data/Nat/Factorization/Basic.lean b/Mathlib/Data/Nat/Factorization/Basic.lean index ca5a516231dff..a56075aac773f 100644 --- a/Mathlib/Data/Nat/Factorization/Basic.lean +++ b/Mathlib/Data/Nat/Factorization/Basic.lean @@ -284,7 +284,7 @@ theorem dvd_iff_div_factorization_eq_tsub {d n : ℕ} (hd : d ≠ 0) (hdn : d d ∣ n ↔ (n / d).factorization = n.factorization - d.factorization := by refine ⟨factorization_div, ?_⟩ rcases eq_or_lt_of_le hdn with (rfl | hd_lt_n); · simp - have h1 : n / d ≠ 0 := fun H => Nat.lt_asymm hd_lt_n ((Nat.div_eq_zero_iff hd.bot_lt).mp H) + have h1 : n / d ≠ 0 := by simp [*] intro h rw [dvd_iff_le_div_mul n d] by_contra h2 diff --git a/Mathlib/Data/Nat/Squarefree.lean b/Mathlib/Data/Nat/Squarefree.lean index 35e092c930493..b1463cae82ce5 100644 --- a/Mathlib/Data/Nat/Squarefree.lean +++ b/Mathlib/Data/Nat/Squarefree.lean @@ -387,8 +387,7 @@ lemma primeFactors_prod (hs : ∀ p ∈ s, p.Prime) : primeFactors (∏ p ∈ s, lemma primeFactors_div_gcd (hm : Squarefree m) (hn : n ≠ 0) : primeFactors (m / m.gcd n) = primeFactors m \ primeFactors n := by ext p - have : m / m.gcd n ≠ 0 := - (Nat.div_ne_zero_iff <| gcd_ne_zero_right hn).2 <| gcd_le_left _ hm.ne_zero.bot_lt + have : m / m.gcd n ≠ 0 := by simp [gcd_ne_zero_right hn, gcd_le_left _ hm.ne_zero.bot_lt] simp only [mem_primeFactors, ne_eq, this, not_false_eq_true, and_true, not_and, mem_sdiff, hm.ne_zero, hn, dvd_div_iff_mul_dvd (gcd_dvd_left _ _)] refine ⟨fun hp ↦ ⟨⟨hp.1, dvd_of_mul_left_dvd hp.2⟩, fun _ hpn ↦ hp.1.not_unit <| hm _ <| diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index 43e41f0704f39..37dffbc1303eb 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -1575,6 +1575,6 @@ def Nat.residueClassesEquiv (N : ℕ) [NeZero N] : ℕ ≃ ZMod N × ℕ where · simp only [add_comm p.1.val, cast_add, cast_mul, natCast_self, zero_mul, natCast_val, cast_id', id_eq, zero_add] · simp only [add_comm p.1.val, mul_add_div (NeZero.pos _), - (Nat.div_eq_zero_iff <| (NeZero.pos _)).2 p.1.val_lt, add_zero] + (Nat.div_eq_zero_iff).2 <| .inr p.1.val_lt, add_zero] set_option linter.style.longFile 1700 From 410afe3f58cac6517a938277a93f6a34e554e18c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sat, 7 Dec 2024 13:24:37 +0000 Subject: [PATCH 565/829] feat(Order/Cofinal): more lemmas on cofinal sets (#19680) --- Mathlib/Order/Cofinal.lean | 26 ++++++++++++++++++++++---- 1 file changed, 22 insertions(+), 4 deletions(-) diff --git a/Mathlib/Order/Cofinal.lean b/Mathlib/Order/Cofinal.lean index 572a4ffc3cedb..532481a8f4cbe 100644 --- a/Mathlib/Order/Cofinal.lean +++ b/Mathlib/Order/Cofinal.lean @@ -3,9 +3,7 @@ Copyright (c) 2024 Violeta Hernández Palacios. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Violeta Hernández Palacios -/ -import Mathlib.Data.Set.Basic -import Mathlib.Tactic.Contrapose -import Mathlib.Order.Bounds.Defs +import Mathlib.Order.GaloisConnection /-! # Cofinal sets @@ -21,7 +19,7 @@ For the cofinality of a set as a cardinal, see `Mathlib.SetTheory.Cardinal.Cofin - Deprecate `Order.Cofinal` in favor of this predicate. -/ -variable {α : Type*} +variable {α β : Type*} section LE variable [LE α] @@ -60,6 +58,16 @@ theorem IsCofinal.trans {s : Set α} {t : Set s} (hs : IsCofinal s) (ht : IsCofi obtain ⟨c, hc, hc'⟩ := ht ⟨b, hb⟩ exact ⟨c, Set.mem_image_of_mem _ hc, hb'.trans hc'⟩ +theorem GaloisConnection.map_cofinal [Preorder β] {f : β → α} {g : α → β} + (h : GaloisConnection f g) {s : Set α} (hs : IsCofinal s) : IsCofinal (g '' s) := by + intro a + obtain ⟨b, hb, hb'⟩ := hs (f a) + exact ⟨g b, Set.mem_image_of_mem _ hb, h.le_iff_le.1 hb'⟩ + +theorem OrderIso.map_cofinal [Preorder β] (e : α ≃o β) {s : Set α} (hs : IsCofinal s) : + IsCofinal (e '' s) := + e.symm.to_galoisConnection.map_cofinal hs + end Preorder section PartialOrder @@ -105,4 +113,14 @@ theorem not_isCofinal_iff_bddAbove [NoMaxOrder α] {s : Set α} : ¬ IsCofinal s theorem not_bddAbove_iff_isCofinal [NoMaxOrder α] {s : Set α} : ¬ BddAbove s ↔ IsCofinal s := not_iff_comm.1 not_isCofinal_iff_bddAbove +/-- The set of "records" (the smallest inputs yielding the highest values) with respect to a +well-ordering of `α` is a cofinal set. -/ +theorem isCofinal_setOf_imp_lt (r : α → α → Prop) [h : IsWellFounded α r] : + IsCofinal { a | ∀ b, r b a → b < a } := by + intro a + obtain ⟨b, hb, hb'⟩ := h.wf.has_min (Set.Ici a) Set.nonempty_Ici + refine ⟨b, fun c hc ↦ ?_, hb⟩ + by_contra! hc' + exact hb' c (hb.trans hc') hc + end LinearOrder From 4fb35bfaca8797d8f5c996fb37f4a8b58bbdc506 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sat, 7 Dec 2024 13:24:38 +0000 Subject: [PATCH 566/829] feat(Order/Basic): `PartialOrder.ext_lt` (#19724) Two partial / linear orders with the same `<` relation are the same. --- Mathlib/Order/Basic.lean | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/Mathlib/Order/Basic.lean b/Mathlib/Order/Basic.lean index 8820922f5b0b4..8e76af6913541 100644 --- a/Mathlib/Order/Basic.lean +++ b/Mathlib/Order/Basic.lean @@ -529,11 +529,20 @@ theorem PartialOrder.ext {A B : PartialOrder α} ext x y exact H x y +theorem PartialOrder.ext_lt {A B : PartialOrder α} + (H : ∀ x y : α, (haveI := A; x < y) ↔ x < y) : A = B := by + ext x y + rw [le_iff_lt_or_eq, @le_iff_lt_or_eq _ A, H] + theorem LinearOrder.ext {A B : LinearOrder α} (H : ∀ x y : α, (haveI := A; x ≤ y) ↔ x ≤ y) : A = B := by ext x y exact H x y +theorem LinearOrder.ext_lt {A B : LinearOrder α} + (H : ∀ x y : α, (haveI := A; x < y) ↔ x < y) : A = B := + LinearOrder.toPartialOrder_injective (PartialOrder.ext_lt H) + /-- Given a relation `R` on `β` and a function `f : α → β`, the preimage relation on `α` is defined by `x ≤ y ↔ f x ≤ f y`. It is the unique relation on `α` making `f` a `RelEmbedding` (assuming `f` is injective). -/ From 6b4fc18b9d4dc8eaae6082bc368489d2998da30f Mon Sep 17 00:00:00 2001 From: Daniel Weber Date: Sat, 7 Dec 2024 13:24:40 +0000 Subject: [PATCH 567/829] =?UTF-8?q?refactor(Order/Basic):=20Make=20`Linear?= =?UTF-8?q?Order=20=CE=B1=20=3D=20LinearOrder=20=CE=B1=E1=B5=92=E1=B5=88?= =?UTF-8?q?=E1=B5=92=E1=B5=88`=20defeq=20(#19776)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/Order/Basic.lean | 17 ++++++++++++++--- 1 file changed, 14 insertions(+), 3 deletions(-) diff --git a/Mathlib/Order/Basic.lean b/Mathlib/Order/Basic.lean index 8e76af6913541..8e5a610e2930f 100644 --- a/Mathlib/Order/Basic.lean +++ b/Mathlib/Order/Basic.lean @@ -698,6 +698,9 @@ instance (α : Type*) [LE α] : LE αᵒᵈ := instance (α : Type*) [LT α] : LT αᵒᵈ := ⟨fun x y : α ↦ y < x⟩ +instance instOrd (α : Type*) [Ord α] : Ord αᵒᵈ where + compare := fun (a b : α) ↦ compare b a + instance instSup (α : Type*) [Min α] : Max αᵒᵈ := ⟨((· ⊓ ·) : α → α → α)⟩ @@ -715,6 +718,7 @@ instance instPartialOrder (α : Type*) [PartialOrder α] : PartialOrder αᵒᵈ instance instLinearOrder (α : Type*) [LinearOrder α] : LinearOrder αᵒᵈ where __ := inferInstanceAs (PartialOrder αᵒᵈ) + __ := inferInstanceAs (Ord αᵒᵈ) le_total := fun a b : α ↦ le_total b a max := fun a b ↦ (min a b : α) min := fun a b ↦ (max a b : α) @@ -722,6 +726,10 @@ instance instLinearOrder (α : Type*) [LinearOrder α] : LinearOrder αᵒᵈ wh max_def := fun a b ↦ show (min .. : α) = _ by rw [min_comm, min_def]; rfl decidableLE := (inferInstance : DecidableRel (fun a b : α ↦ b ≤ a)) decidableLT := (inferInstance : DecidableRel (fun a b : α ↦ b < a)) + decidableEq := (inferInstance : DecidableEq α) + compare_eq_compareOfLessAndEq a b := by + simp only [compare, LinearOrder.compare_eq_compareOfLessAndEq, compareOfLessAndEq, eq_comm] + rfl /-- The opposite linear order to a given linear order -/ def _root_.LinearOrder.swap (α : Type*) (_ : LinearOrder α) : LinearOrder α := @@ -729,16 +737,19 @@ def _root_.LinearOrder.swap (α : Type*) (_ : LinearOrder α) : LinearOrder α : instance : ∀ [Inhabited α], Inhabited αᵒᵈ := fun [x : Inhabited α] => x +theorem Ord.dual_dual (α : Type*) [H : Ord α] : OrderDual.instOrd αᵒᵈ = H := + rfl + theorem Preorder.dual_dual (α : Type*) [H : Preorder α] : OrderDual.instPreorder αᵒᵈ = H := - Preorder.ext fun _ _ ↦ Iff.rfl + rfl theorem instPartialOrder.dual_dual (α : Type*) [H : PartialOrder α] : OrderDual.instPartialOrder αᵒᵈ = H := - PartialOrder.ext fun _ _ ↦ Iff.rfl + rfl theorem instLinearOrder.dual_dual (α : Type*) [H : LinearOrder α] : OrderDual.instLinearOrder αᵒᵈ = H := - LinearOrder.ext fun _ _ ↦ Iff.rfl + rfl end OrderDual From a988e8b369e20ae90b4998923e4efcd732a1da62 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Sat, 7 Dec 2024 14:05:32 +0000 Subject: [PATCH 568/829] chore(NumberTheory/Numberfield/*): rename ringEquiv_complex_of_isComplex -> ringEquivComplexOfIsComplex (#19659) --- .../NumberTheory/NumberField/AdeleRing.lean | 10 ++--- .../NumberTheory/NumberField/Completion.lean | 45 ++++++++++++------- 2 files changed, 35 insertions(+), 20 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/AdeleRing.lean b/Mathlib/NumberTheory/NumberField/AdeleRing.lean index fbf792b8edcbb..1ee25ce15fd5f 100644 --- a/Mathlib/NumberTheory/NumberField/AdeleRing.lean +++ b/Mathlib/NumberTheory/NumberField/AdeleRing.lean @@ -84,9 +84,9 @@ abbrev ringEquiv_mixedSpace : (RingEquiv.piEquivPiSubtypeProd (fun (v : InfinitePlace K) => IsReal v) (fun (v : InfinitePlace K) => v.Completion)) (RingEquiv.prodCongr - (RingEquiv.piCongrRight (fun ⟨_, hv⟩ => Completion.ringEquiv_real_of_isReal hv)) + (RingEquiv.piCongrRight (fun ⟨_, hv⟩ => Completion.ringEquivRealOfIsReal hv)) (RingEquiv.trans - (RingEquiv.piCongrRight (fun v => Completion.ringEquiv_complex_of_isComplex + (RingEquiv.piCongrRight (fun v => Completion.ringEquivComplexOfIsComplex ((not_isReal_iff_isComplex.1 v.2)))) (RingEquiv.piCongrLeft (fun _ => ℂ) <| Equiv.subtypeEquivRight (fun _ => not_isReal_iff_isComplex)))) @@ -94,7 +94,7 @@ abbrev ringEquiv_mixedSpace : @[simp] theorem ringEquiv_mixedSpace_apply (x : InfiniteAdeleRing K) : ringEquiv_mixedSpace K x = - (fun (v : {w : InfinitePlace K // IsReal w}) => extensionEmbedding_of_isReal v.2 (x v), + (fun (v : {w : InfinitePlace K // IsReal w}) => extensionEmbeddingOfIsReal v.2 (x v), fun (v : {w : InfinitePlace K // IsComplex w}) => extensionEmbedding v.1 (x v)) := rfl /-- Transfers the embedding of `x ↦ (x)ᵥ` of the number field `K` into its infinite adele @@ -103,8 +103,8 @@ ring to the mixed embedding `x ↦ (φᵢ(x))ᵢ` of `K` into the space `ℝ ^ r theorem mixedEmbedding_eq_algebraMap_comp {x : K} : mixedEmbedding K x = ringEquiv_mixedSpace K (algebraMap K _ x) := by ext v <;> simp only [ringEquiv_mixedSpace_apply, algebraMap_apply, - ringEquiv_real_of_isReal, ringEquiv_complex_of_isComplex, extensionEmbedding, - extensionEmbedding_of_isReal, extensionEmbedding_of_comp, RingEquiv.coe_ofBijective, + ringEquivRealOfIsReal, ringEquivComplexOfIsComplex, extensionEmbedding, + extensionEmbeddingOfIsReal, extensionEmbedding_of_comp, RingEquiv.coe_ofBijective, RingHom.coe_mk, MonoidHom.coe_mk, OneHom.coe_mk, UniformSpace.Completion.extensionHom] · rw [UniformSpace.Completion.extension_coe (WithAbs.isUniformInducing_of_comp <| v.1.norm_embedding_of_isReal v.2).uniformContinuous x] diff --git a/Mathlib/NumberTheory/NumberField/Completion.lean b/Mathlib/NumberTheory/NumberField/Completion.lean index 155d6e66f7450..0d23896da11ed 100644 --- a/Mathlib/NumberTheory/NumberField/Completion.lean +++ b/Mathlib/NumberTheory/NumberField/Completion.lean @@ -29,12 +29,12 @@ field at an infinite place is then derived in this file, as `InfinitePlace` is a place. - `NumberField.InfinitePlace.Completion.extensionEmbedding` : the embedding `v.embedding : K →+* ℂ` extended to `v.Completion →+* ℂ`. - - `NumberField.InfinitePlace.Completion.extensionEmbedding_of_isReal` : if the infinite place `v` + - `NumberField.InfinitePlace.Completion.extensionEmbeddingOfIsReal` : if the infinite place `v` is real, then this extends the embedding `v.embedding_of_isReal : K →+* ℝ` to `v.Completion →+* ℝ`. - `NumberField.InfinitePlace.Completion.equiv_real_of_isReal` : the ring isomorphism `v.Completion ≃+* ℝ` when `v` is a real infinite place; the forward direction of this is - `extensionEmbedding_of_isReal`. + `extensionEmbeddingOfIsReal`. - `NumberField.InfinitePlace.Completion.equiv_complex_of_isComplex` : the ring isomorphism `v.Completion ≃+* ℂ` when `v` is a complex infinite place; the forward direction of this is `extensionEmbedding`. @@ -99,16 +99,19 @@ instance locallyCompactSpace : LocallyCompactSpace (v.Completion) := def extensionEmbedding : v.Completion →+* ℂ := extensionEmbedding_of_comp v.norm_embedding_eq /-- The embedding `K →+* ℝ` associated to a real infinite place extended to `v.Completion →+* ℝ`. -/ -def extensionEmbedding_of_isReal {v : InfinitePlace K} (hv : IsReal v) : v.Completion →+* ℝ := +def extensionEmbeddingOfIsReal {v : InfinitePlace K} (hv : IsReal v) : v.Completion →+* ℝ := extensionEmbedding_of_comp <| v.norm_embedding_of_isReal hv +@[deprecated (since := "2024-12-07")] +noncomputable alias extensionEmbedding_of_isReal := extensionEmbeddingOfIsReal + @[simp] theorem extensionEmbedding_coe (x : K) : extensionEmbedding v x = v.embedding x := extensionEmbedding_of_comp_coe v.norm_embedding_eq x @[simp] theorem extensionEmbedding_of_isReal_coe {v : InfinitePlace K} (hv : IsReal v) (x : K) : - extensionEmbedding_of_isReal hv x = embedding_of_isReal hv x := + extensionEmbeddingOfIsReal hv x = embedding_of_isReal hv x := extensionEmbedding_of_comp_coe (v.norm_embedding_of_isReal hv) x /-- The embedding `v.Completion →+* ℂ` is an isometry. -/ @@ -117,7 +120,7 @@ theorem isometry_extensionEmbedding : Isometry (extensionEmbedding v) := /-- The embedding `v.Completion →+* ℝ` at a real infinite palce is an isometry. -/ theorem isometry_extensionEmbedding_of_isReal {v : InfinitePlace K} (hv : IsReal v) : - Isometry (extensionEmbedding_of_isReal hv) := + Isometry (extensionEmbeddingOfIsReal hv) := Isometry.of_dist_eq (extensionEmbedding_dist_eq_of_comp <| v.norm_embedding_of_isReal hv) /-- The embedding `v.Completion →+* ℂ` has closed image inside `ℂ`. -/ @@ -127,7 +130,7 @@ theorem isClosed_image_extensionEmbedding : IsClosed (Set.range (extensionEmbedd /-- The embedding `v.Completion →+* ℝ` associated to a real infinite place has closed image inside `ℝ`. -/ theorem isClosed_image_extensionEmbedding_of_isReal {v : InfinitePlace K} (hv : IsReal v) : - IsClosed (Set.range (extensionEmbedding_of_isReal hv)) := + IsClosed (Set.range (extensionEmbeddingOfIsReal hv)) := (isClosedEmbedding_extensionEmbedding_of_comp <| v.norm_embedding_of_isReal hv).isClosed_range theorem subfield_ne_real_of_isComplex {v : InfinitePlace K} (hv : IsComplex v) : @@ -152,35 +155,47 @@ theorem bijective_extensionEmbedding_of_isComplex {v : InfinitePlace K} (hv : Is /-- The ring isomorphism `v.Completion ≃+* ℂ`, when `v` is complex, given by the bijection `v.Completion →+* ℂ`. -/ -def ringEquiv_complex_of_isComplex {v : InfinitePlace K} (hv : IsComplex v) : +def ringEquivComplexOfIsComplex {v : InfinitePlace K} (hv : IsComplex v) : v.Completion ≃+* ℂ := RingEquiv.ofBijective _ (bijective_extensionEmbedding_of_isComplex hv) +@[deprecated (since := "2024-12-07")] +noncomputable alias ringEquiv_complex_of_isComplex := ringEquivComplexOfIsComplex + /-- If the infinite place `v` is complex, then `v.Completion` is isometric to `ℂ`. -/ -def isometryEquiv_complex_of_isComplex {v : InfinitePlace K} (hv : IsComplex v) : +def isometryEquivComplexOfIsComplex {v : InfinitePlace K} (hv : IsComplex v) : v.Completion ≃ᵢ ℂ where - toEquiv := ringEquiv_complex_of_isComplex hv + toEquiv := ringEquivComplexOfIsComplex hv isometry_toFun := isometry_extensionEmbedding v +@[deprecated (since := "2024-12-07")] +noncomputable alias isometryEquiv_complex_of_isComplex := isometryEquivComplexOfIsComplex + /-- If `v` is a real infinite place, then the embedding `v.Completion →+* ℝ` is surjective. -/ theorem surjective_extensionEmbedding_of_isReal {v : InfinitePlace K} (hv : IsReal v) : - Function.Surjective (extensionEmbedding_of_isReal hv) := by + Function.Surjective (extensionEmbeddingOfIsReal hv) := by rw [← RingHom.fieldRange_eq_top_iff, ← Real.subfield_eq_of_closed] exact isClosed_image_extensionEmbedding_of_isReal hv /-- If `v` is a real infinite place, then the embedding `v.Completion →+* ℝ` is bijective. -/ theorem bijective_extensionEmbedding_of_isReal {v : InfinitePlace K} (hv : IsReal v) : - Function.Bijective (extensionEmbedding_of_isReal hv) := - ⟨(extensionEmbedding_of_isReal hv).injective, surjective_extensionEmbedding_of_isReal hv⟩ + Function.Bijective (extensionEmbeddingOfIsReal hv) := + ⟨(extensionEmbeddingOfIsReal hv).injective, surjective_extensionEmbedding_of_isReal hv⟩ /-- The ring isomorphism `v.Completion ≃+* ℝ`, when `v` is real, given by the bijection `v.Completion →+* ℝ`. -/ -def ringEquiv_real_of_isReal {v : InfinitePlace K} (hv : IsReal v) : v.Completion ≃+* ℝ := +def ringEquivRealOfIsReal {v : InfinitePlace K} (hv : IsReal v) : v.Completion ≃+* ℝ := RingEquiv.ofBijective _ (bijective_extensionEmbedding_of_isReal hv) +@[deprecated (since := "2024-12-07")] +noncomputable alias ringEquiv_real_of_isReal := ringEquivRealOfIsReal + /-- If the infinite place `v` is real, then `v.Completion` is isometric to `ℝ`. -/ -def isometryEquiv_real_of_isReal {v : InfinitePlace K} (hv : IsReal v) : v.Completion ≃ᵢ ℝ where - toEquiv := ringEquiv_real_of_isReal hv +def isometryEquivRealOfIsReal {v : InfinitePlace K} (hv : IsReal v) : v.Completion ≃ᵢ ℝ where + toEquiv := ringEquivRealOfIsReal hv isometry_toFun := isometry_extensionEmbedding_of_isReal hv +@[deprecated (since := "2024-12-07")] +noncomputable alias isometryEquiv_real_of_isReal := isometryEquivRealOfIsReal + end NumberField.InfinitePlace.Completion From 9a3b5ca02eeb6b8821b1adeed6d540da55cbf469 Mon Sep 17 00:00:00 2001 From: leanprover-community-bot-assistant Date: Sun, 8 Dec 2024 00:16:25 +0000 Subject: [PATCH 569/829] chore(scripts): update nolints.json (#19791) I am happy to remove some nolints for you! --- scripts/nolints.json | 1 - 1 file changed, 1 deletion(-) diff --git a/scripts/nolints.json b/scripts/nolints.json index 4f5d00a571a78..58730127f796d 100644 --- a/scripts/nolints.json +++ b/scripts/nolints.json @@ -352,7 +352,6 @@ ["docBlame", "VectorPrebundle.pretrivializationAt"], ["docBlame", "VectorPrebundle.pretrivializationAtlas"], ["docBlame", "WithIdeal.i"], - ["docBlame", "WithZeroTopology.instHasContinuousInv₀"], ["docBlame", "WriterT.adapt"], ["docBlame", "WriterT.callCC"], ["docBlame", "WriterT.callCC'"], From 19b191ae82eabc6ade3c07cc325998bd8bee478a Mon Sep 17 00:00:00 2001 From: Mitchell Lee Date: Sun, 8 Dec 2024 02:15:37 +0000 Subject: [PATCH 570/829] doc(Combinatorics/Young/YoungDiagram): Change `Ssyt` to `SemistandardYoungTableau` (#19773) --- Mathlib/Combinatorics/Young/SemistandardTableau.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Combinatorics/Young/SemistandardTableau.lean b/Mathlib/Combinatorics/Young/SemistandardTableau.lean index fac50829baf05..52168fe27c5c2 100644 --- a/Mathlib/Combinatorics/Young/SemistandardTableau.lean +++ b/Mathlib/Combinatorics/Young/SemistandardTableau.lean @@ -29,8 +29,8 @@ for all pairs `(i, j) ∉ μ` and to satisfy the row-weak and column-strict cond - `SemistandardYoungTableau (μ : YoungDiagram)`: semistandard Young tableaux of shape `μ`. There is a `coe` instance such that `T i j` is value of the `(i, j)` entry of the semistandard Young tableau `T`. -- `Ssyt.highestWeight (μ : YoungDiagram)`: the semistandard Young tableau whose `i`th row - consists entirely of `i`s, for each `i`. +- `SemistandardYoungTableau.highestWeight (μ : YoungDiagram)`: the semistandard Young tableau whose + `i`th row consists entirely of `i`s, for each `i`. ## Tags From 31530f31d529f37cde81d086fe10b68e9eb1ee2f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 8 Dec 2024 07:41:29 +0000 Subject: [PATCH 571/829] chore(AdicCompletion): generalise algebra instance to separate rings (#19466) In the existing instance `Algebra R (AdicCompletion I R)`, `R` appears three times: On the left, on the right, and in `I : Ideal R`. The left occurrence can be generalised to a ring `S` such that `R` is a `S`-algebra. Closes ImperialCollegeLondon/FLT/issues/230. From FLT --- .../RingTheory/AdicCompletion/Algebra.lean | 26 ++--- .../AdicCompletion/AsTensorProduct.lean | 2 +- Mathlib/RingTheory/AdicCompletion/Basic.lean | 104 +++++++++++------- .../AdicCompletion/Functoriality.lean | 2 +- 4 files changed, 76 insertions(+), 58 deletions(-) diff --git a/Mathlib/RingTheory/AdicCompletion/Algebra.lean b/Mathlib/RingTheory/AdicCompletion/Algebra.lean index 6f7c6168e2c6c..9c8babe6957b6 100644 --- a/Mathlib/RingTheory/AdicCompletion/Algebra.lean +++ b/Mathlib/RingTheory/AdicCompletion/Algebra.lean @@ -25,7 +25,7 @@ suppress_compilation open Submodule -variable {R : Type*} [CommRing R] (I : Ideal R) +variable {R S : Type*} [CommRing R] [CommRing S] (I : Ideal R) variable {M : Type*} [AddCommGroup M] [Module R M] namespace AdicCompletion @@ -87,8 +87,10 @@ instance : CommRing (AdicCompletion I R) := (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ ↦ rfl) -instance : Algebra R (AdicCompletion I R) where - toFun r := ⟨algebraMap R (∀ n, R ⧸ (I ^ n • ⊤ : Ideal R)) r, by simp⟩ +instance [Algebra S R] : Algebra S (AdicCompletion I R) where + toFun r := ⟨algebraMap S (∀ n, R ⧸ (I ^ n • ⊤ : Ideal R)) r, by + simp [-Ideal.Quotient.mk_algebraMap, + IsScalarTower.algebraMap_apply S R (R ⧸ (I ^ _ • ⊤ : Ideal R))]⟩ map_one' := Subtype.ext <| map_one _ map_mul' x y := Subtype.ext <| map_mul _ x y map_zero' := Subtype.ext <| map_zero _ @@ -252,23 +254,15 @@ instance module : Module (AdicCompletion I R) (AdicCompletion I M) where mul_smul r s x := by ext n simp only [smul_eval, val_mul, mul_smul] - smul_zero r := by - ext n - rw [smul_eval, val_zero, smul_zero] - smul_add r x y := by - ext n - simp only [smul_eval, val_add, smul_add] - add_smul r s x := by - ext n - simp only [coe_eval, smul_eval, map_add, add_smul, val_add] - zero_smul x := by - ext n - simp only [smul_eval, _root_.map_zero, zero_smul, val_zero] + smul_zero r := by ext n; simp + smul_add r x y := by ext n; simp + add_smul r s x := by ext n; simp [val_smul, add_smul] + zero_smul x := by ext n; simp instance : IsScalarTower R (AdicCompletion I R) (AdicCompletion I M) where smul_assoc r s x := by ext n - rw [smul_eval, val_smul, val_smul, smul_eval, smul_assoc] + rw [smul_eval, val_smul_apply, val_smul_apply, smul_eval, smul_assoc] /-- A priori `AdicCompletion I R` has two `AdicCompletion I R`-module instances. Both agree definitionally. -/ diff --git a/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean b/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean index bd80f8783d425..af49d74e4ff2a 100644 --- a/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean +++ b/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean @@ -101,7 +101,7 @@ private lemma piEquivOfFintype_comp_ofTensorProduct_eq : ext i j k suffices h : (if j = i then 1 else 0) = (if j = i then 1 else 0 : AdicCompletion I R).val k by simpa [Pi.single_apply, -smul_eq_mul, -Algebra.id.smul_eq_mul] - split <;> simp only [smul_eq_mul, val_zero, val_one] + split <;> simp private lemma ofTensorProduct_eq : ofTensorProduct I (ι → R) = (piEquivOfFintype I (ι := ι) (fun _ : ι ↦ R)).symm.toLinearMap ∘ₗ diff --git a/Mathlib/RingTheory/AdicCompletion/Basic.lean b/Mathlib/RingTheory/AdicCompletion/Basic.lean index e1dbd4e2cb74f..62028936b50db 100644 --- a/Mathlib/RingTheory/AdicCompletion/Basic.lean +++ b/Mathlib/RingTheory/AdicCompletion/Basic.lean @@ -29,7 +29,7 @@ suppress_compilation open Submodule -variable {R : Type*} [CommRing R] (I : Ideal R) +variable {R S T : Type*} [CommRing R] (I : Ideal R) variable (M : Type*) [AddCommGroup M] [Module R M] variable {N : Type*} [AddCommGroup N] [Module R N] @@ -202,24 +202,62 @@ instance : Neg (AdicCompletion I M) where instance : Sub (AdicCompletion I M) where sub x y := ⟨x.val - y.val, by simp [x.property, y.property]⟩ -instance : SMul ℕ (AdicCompletion I M) where - smul n x := ⟨n • x.val, by simp [x.property]⟩ +instance instSMul [SMul S R] [SMul S M] [IsScalarTower S R M] : SMul S (AdicCompletion I M) where + smul r x := ⟨r • x.val, by simp [x.property]⟩ + +@[simp, norm_cast] lemma val_zero : (0 : AdicCompletion I M).val = 0 := rfl + +lemma val_zero_apply (n : ℕ) : (0 : AdicCompletion I M).val n = 0 := rfl + +variable {I M} + +@[simp, norm_cast] lemma val_add (f g : AdicCompletion I M) : (f + g).val = f.val + g.val := rfl +@[simp, norm_cast] lemma val_sub (f g : AdicCompletion I M) : (f - g).val = f.val - g.val := rfl +@[simp, norm_cast] lemma val_neg (f : AdicCompletion I M) : (-f).val = -f.val := rfl + +lemma val_add_apply (f g : AdicCompletion I M) (n : ℕ) : (f + g).val n = f.val n + g.val n := rfl +lemma val_sub_apply (f g : AdicCompletion I M) (n : ℕ) : (f - g).val n = f.val n - g.val n := rfl +lemma val_neg_apply (f : AdicCompletion I M) (n : ℕ) : (-f).val n = -f.val n := rfl + +/- No `simp` attribute, since it causes `simp` unification timeouts when considering +the `Module (AdicCompletion I R) (AdicCompletion I M)` instance (see `AdicCompletion/Algebra`). -/ +@[norm_cast] +lemma val_smul [SMul S R] [SMul S M] [IsScalarTower S R M] (s : S) (f : AdicCompletion I M) : + (s • f).val = s • f.val := rfl + +lemma val_smul_apply [SMul S R] [SMul S M] [IsScalarTower S R M] (s : S) (f : AdicCompletion I M) + (n : ℕ) : (s • f).val n = s • f.val n := rfl + +@[ext] +lemma ext {x y : AdicCompletion I M} (h : ∀ n, x.val n = y.val n) : x = y := Subtype.eq <| funext h -instance : SMul ℤ (AdicCompletion I M) where - smul n x := ⟨n • x.val, by simp [x.property]⟩ +variable (I M) instance : AddCommGroup (AdicCompletion I M) := let f : AdicCompletion I M → ∀ n, M ⧸ (I ^ n • ⊤ : Submodule R M) := Subtype.val - Subtype.val_injective.addCommGroup f rfl (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ _ ↦ rfl) - (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) + Subtype.val_injective.addCommGroup f rfl val_add val_neg val_sub (fun _ _ ↦ val_smul ..) + (fun _ _ ↦ val_smul ..) -instance : SMul R (AdicCompletion I M) where - smul r x := ⟨r • x.val, by simp [x.property]⟩ - -instance : Module R (AdicCompletion I M) := +instance [Semiring S] [SMul S R] [Module S M] [IsScalarTower S R M] : + Module S (AdicCompletion I M) := let f : AdicCompletion I M →+ ∀ n, M ⧸ (I ^ n • ⊤ : Submodule R M) := { toFun := Subtype.val, map_zero' := rfl, map_add' := fun _ _ ↦ rfl } - Subtype.val_injective.module R f (fun _ _ ↦ rfl) + Subtype.val_injective.module S f val_smul + +instance instIsScalarTower [SMul S T] [SMul S R] [SMul T R] [SMul S M] [SMul T M] + [IsScalarTower S R M] [IsScalarTower T R M] [IsScalarTower S T M] : + IsScalarTower S T (AdicCompletion I M) where + smul_assoc s t f := by ext; simp [val_smul] + +instance instSMulCommClass [SMul S R] [SMul T R] [SMul S M] [SMul T M] + [IsScalarTower S R M] [IsScalarTower T R M] [SMulCommClass S T M] : + SMulCommClass S T (AdicCompletion I M) where + smul_comm s t f := by ext; simp [val_smul, smul_comm] + +instance instIsCentralScalar [SMul S R] [SMul Sᵐᵒᵖ R] [SMul S M] [SMul Sᵐᵒᵖ M] + [IsScalarTower S R M] [IsScalarTower Sᵐᵒᵖ R M] [IsCentralScalar S M] : + IsCentralScalar S (AdicCompletion I M) where + op_smul_eq_smul s f := by ext; simp [val_smul, op_smul_eq_smul] /-- The canonical inclusion from the completion to the product. -/ @[simps] @@ -228,6 +266,18 @@ def incl : AdicCompletion I M →ₗ[R] (∀ n, M ⧸ (I ^ n • ⊤ : Submodule map_add' _ _ := rfl map_smul' _ _ := rfl +variable {I M} + +@[simp, norm_cast] +lemma val_sum {ι : Type*} (s : Finset ι) (f : ι → AdicCompletion I M) : + (∑ i ∈ s, f i).val = ∑ i ∈ s, (f i).val := by + simp_rw [← funext (incl_apply _ _ _), map_sum] + +lemma val_sum_apply {ι : Type*} (s : Finset ι) (f : ι → AdicCompletion I M) (n : ℕ) : + (∑ i ∈ s, f i).val n = ∑ i ∈ s, (f i).val n := by simp + +variable (I M) + /-- The canonical linear map to the completion. -/ def of : M →ₗ[R] AdicCompletion I M where toFun x := ⟨fun n => mkQ (I ^ n • ⊤ : Submodule R M) x, fun _ => rfl⟩ @@ -266,43 +316,17 @@ theorem eval_surjective (n : ℕ) : Function.Surjective (eval I M n) := fun x theorem range_eval (n : ℕ) : LinearMap.range (eval I M n) = ⊤ := LinearMap.range_eq_top.2 (eval_surjective I M n) -@[simp] -theorem val_zero (n : ℕ) : (0 : AdicCompletion I M).val n = 0 := - rfl - variable {I M} -@[simp] -theorem val_add (n : ℕ) (f g : AdicCompletion I M) : (f + g).val n = f.val n + g.val n := - rfl - -@[simp] -theorem val_sub (n : ℕ) (f g : AdicCompletion I M) : (f - g).val n = f.val n - g.val n := - rfl - -@[simp] -theorem val_sum {α : Type*} (s : Finset α) (f : α → AdicCompletion I M) (n : ℕ) : - (Finset.sum s f).val n = Finset.sum s (fun a ↦ (f a).val n) := by - simp_rw [← incl_apply, map_sum, Finset.sum_apply] - -/- No `simp` attribute, since it causes `simp` unification timeouts when considering -the `AdicCompletion I R` module instance on `AdicCompletion I M` (see `AdicCompletion/Algebra`). -/ -theorem val_smul (n : ℕ) (r : R) (f : AdicCompletion I M) : (r • f).val n = r • f.val n := - rfl - -@[ext] -theorem ext {x y : AdicCompletion I M} (h : ∀ n, x.val n = y.val n) : x = y := - Subtype.eq <| funext h - variable (I M) instance : IsHausdorff I (AdicCompletion I M) where haus' x h := ext fun n ↦ by refine smul_induction_on (SModEq.zero.1 <| h n) (fun r hr x _ ↦ ?_) (fun x y hx hy ↦ ?_) - · simp only [val_smul, val_zero] + · simp only [val_smul_apply, val_zero] exact Quotient.inductionOn' (x.val n) (fun a ↦ SModEq.zero.2 <| smul_mem_smul hr mem_top) - · simp only [val_add, hx, val_zero, hy, add_zero] + · simp only [val_add_apply, hx, val_zero_apply, hy, add_zero] @[simp] theorem transitionMap_mk {m n : ℕ} (hmn : m ≤ n) (x : M) : diff --git a/Mathlib/RingTheory/AdicCompletion/Functoriality.lean b/Mathlib/RingTheory/AdicCompletion/Functoriality.lean index 03b236a454123..4ac8d7b974ccc 100644 --- a/Mathlib/RingTheory/AdicCompletion/Functoriality.lean +++ b/Mathlib/RingTheory/AdicCompletion/Functoriality.lean @@ -302,7 +302,7 @@ theorem sum_comp_sumInv : sum I M ∘ₗ sumInv I M = LinearMap.id := by simp only [LinearMap.coe_comp, Function.comp_apply, LinearMap.id_coe, id_eq, mk_apply_coe, Submodule.mkQ_apply] rw [← DirectSum.sum_univ_of (((sumInv I M) ((AdicCompletion.mk I (⨁ (j : ι), M j)) f)))] - simp only [sumInv_apply, map_mk, map_sum, sum_of, val_sum, mk_apply_coe, + simp only [sumInv_apply, map_mk, map_sum, sum_of, val_sum_apply, mk_apply_coe, AdicCauchySequence.map_apply_coe, Submodule.mkQ_apply] simp only [← Submodule.mkQ_apply, ← map_sum] erw [DirectSum.sum_univ_of] From 294458c5add8df2c106f392d97e53aec900e95f5 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 8 Dec 2024 15:02:00 +0000 Subject: [PATCH 572/829] chore(*): rename some `defs` to `camelCase` (#19795) --- Mathlib/CategoryTheory/Monoidal/Bimon_.lean | 10 ++++++++-- Mathlib/Data/Set/Lattice.lean | 4 +++- Mathlib/LinearAlgebra/FreeProduct/Basic.lean | 19 +++++++++++++------ Mathlib/Tactic/Linter/OldObtain.lean | 7 +++++-- Mathlib/Tactic/Linter/Style.lean | 18 +++++++++++++----- Mathlib/Tactic/MoveAdd.lean | 7 +++++-- 6 files changed, 47 insertions(+), 18 deletions(-) diff --git a/Mathlib/CategoryTheory/Monoidal/Bimon_.lean b/Mathlib/CategoryTheory/Monoidal/Bimon_.lean index 4b3678918fac2..69026eb1f7464 100644 --- a/Mathlib/CategoryTheory/Monoidal/Bimon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Bimon_.lean @@ -157,14 +157,20 @@ def trivial : Bimon_ C := Comon_.trivial (Mon_ C) /-- The bimonoid morphism from the trivial bimonoid to any bimonoid. -/ @[simps] -def trivial_to (A : Bimon_ C) : trivial C ⟶ A := +def trivialTo (A : Bimon_ C) : trivial C ⟶ A := { hom := (default : Mon_.trivial C ⟶ A.X), } +@[deprecated (since := "2024-12-07")] alias trivial_to := trivialTo +@[deprecated (since := "2024-12-07")] alias trivial_to_hom := trivialTo_hom + /-- The bimonoid morphism from any bimonoid to the trivial bimonoid. -/ @[simps!] -def to_trivial (A : Bimon_ C) : A ⟶ trivial C := +def toTrivial (A : Bimon_ C) : A ⟶ trivial C := (default : @Quiver.Hom (Comon_ (Mon_ C)) _ A (Comon_.trivial (Mon_ C))) +@[deprecated (since := "2024-12-07")] alias to_trivial := toTrivial +@[deprecated (since := "2024-12-07")] alias to_trivial_hom := toTrivial_hom + /-! # Additional lemmas -/ variable {C} diff --git a/Mathlib/Data/Set/Lattice.lean b/Mathlib/Data/Set/Lattice.lean index 41676f753e553..b77cd5d86cdec 100644 --- a/Mathlib/Data/Set/Lattice.lean +++ b/Mathlib/Data/Set/Lattice.lean @@ -924,10 +924,12 @@ theorem sUnion_powerset_gc : gc_sSup_Iic /-- `⋃₀` and `𝒫` form a Galois insertion. -/ -def sUnion_powerset_gi : +def sUnionPowersetGI : GaloisInsertion (⋃₀ · : Set (Set α) → Set α) (𝒫 · : Set α → Set (Set α)) := gi_sSup_Iic +@[deprecated (since := "2024-12-07")] alias sUnion_powerset_gi := sUnionPowersetGI + /-- If all sets in a collection are either `∅` or `Set.univ`, then so is their union. -/ theorem sUnion_mem_empty_univ {S : Set (Set α)} (h : S ⊆ {∅, univ}) : ⋃₀ S ∈ ({∅, univ} : Set (Set α)) := by diff --git a/Mathlib/LinearAlgebra/FreeProduct/Basic.lean b/Mathlib/LinearAlgebra/FreeProduct/Basic.lean index bfa28160ab5c3..a6d1297916a19 100644 --- a/Mathlib/LinearAlgebra/FreeProduct/Basic.lean +++ b/Mathlib/LinearAlgebra/FreeProduct/Basic.lean @@ -68,7 +68,7 @@ universe uS uA uB then their quotients are also `R`-equivalent. (Special case of the third isomorphism theorem.)-/ -def algEquiv_quot_algEquiv +def algEquivQuotAlgEquiv {R : Type u} [CommSemiring R] {A B : Type v} [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] (f : A ≃ₐ[R] B) (rel : A → A → Prop) : RingQuot rel ≃ₐ[R] RingQuot (rel on f.symm) := @@ -83,16 +83,20 @@ def algEquiv_quot_algEquiv fun x y h ↦ by apply RingQuot.mkAlgHom_rel; simpa⟩)) (by ext b; simp) (by ext a; simp) +@[deprecated (since := "2024-12-07")] alias algEquiv_quot_algEquiv := algEquivQuotAlgEquiv + /--If two (semi)rings are equivalent and their quotients by a relation `rel` are defined, then their quotients are also equivalent. (Special case of `algEquiv_quot_algEquiv` when `R = ℕ`, which in turn is a special case of the third isomorphism theorem.)-/ -def equiv_quot_equiv {A B : Type v} [Semiring A] [Semiring B] (f : A ≃+* B) (rel : A → A → Prop) : +def equivQuotEquiv {A B : Type v} [Semiring A] [Semiring B] (f : A ≃+* B) (rel : A → A → Prop) : RingQuot rel ≃+* RingQuot (rel on f.symm) := let f_alg : A ≃ₐ[ℕ] B := AlgEquiv.ofRingEquiv (f := f) (fun n ↦ by simp) - algEquiv_quot_algEquiv f_alg rel |>.toRingEquiv + algEquivQuotAlgEquiv f_alg rel |>.toRingEquiv + +@[deprecated (since := "2024-12-07")] alias equiv_quot_equiv := equivQuotEquiv end RingQuot @@ -143,11 +147,14 @@ theorem rel_id (i : I) : rel R A (ι R <| lof R I A i 1) 1 := rel.id @[reducible] def _root_.LinearAlgebra.FreeProduct := RingQuot <| FreeProduct.rel R A /--The free product of the collection of `R`-algebras `A i`, as a quotient of `PowerAlgebra R A`-/ -@[reducible] def _root_.LinearAlgebra.FreeProduct_ofPowers := RingQuot <| FreeProduct.rel' R A +@[reducible] def _root_.LinearAlgebra.FreeProductOfPowers := RingQuot <| FreeProduct.rel' R A + +@[deprecated (since := "2024-12-07")] +alias _root_.LinearAlgebra.FreeProduct_ofPowers := LinearAlgebra.FreeProductOfPowers /--The `R`-algebra equivalence relating `FreeProduct` and `FreeProduct_ofPowers`-/ -noncomputable def equivPowerAlgebra : FreeProduct_ofPowers R A ≃ₐ[R] FreeProduct R A := - RingQuot.algEquiv_quot_algEquiv +noncomputable def equivPowerAlgebra : FreeProductOfPowers R A ≃ₐ[R] FreeProduct R A := + RingQuot.algEquivQuotAlgEquiv (FreeProduct.powerAlgebra_equiv_freeAlgebra R A |>.symm) (FreeProduct.rel R A) |>.symm diff --git a/Mathlib/Tactic/Linter/OldObtain.lean b/Mathlib/Tactic/Linter/OldObtain.lean index 34997e37d0fc6..f9d77e63e79e8 100644 --- a/Mathlib/Tactic/Linter/OldObtain.lean +++ b/Mathlib/Tactic/Linter/OldObtain.lean @@ -54,12 +54,15 @@ open Lean Elab namespace Mathlib.Linter.Style /-- Whether a syntax element is an `obtain` tactic call without a provided proof. -/ -def is_obtain_without_proof : Syntax → Bool +def isObtainWithoutProof : Syntax → Bool -- Using the `obtain` tactic without a proof requires proving a type; -- a pattern is optional. | `(tactic|obtain : $_type) | `(tactic|obtain $_pat : $_type) => true | _ => false +/-- Deprecated alias for `Mathlib.Linter.Style.isObtainWithoutProof`. -/ +@[deprecated isObtainWithoutProof (since := "2024-12-07")] +def is_obtain_without_proof := @isObtainWithoutProof /-- The `oldObtain` linter emits a warning upon uses of the "stream-of-conciousness" variants of the `obtain` tactic, i.e. with the proof postponed. -/ @@ -74,7 +77,7 @@ def oldObtainLinter : Linter where run := withSetOptionIn fun stx => do return if (← MonadState.get).messages.hasErrors then return - if let some head := stx.find? is_obtain_without_proof then + if let some head := stx.find? isObtainWithoutProof then Linter.logLint linter.oldObtain head m!"Please remove stream-of-conciousness `obtain` syntax" initialize addLinter oldObtainLinter diff --git a/Mathlib/Tactic/Linter/Style.lean b/Mathlib/Tactic/Linter/Style.lean index 65272ba2d6d0a..bbe2f16f58a98 100644 --- a/Mathlib/Tactic/Linter/Style.lean +++ b/Mathlib/Tactic/Linter/Style.lean @@ -52,16 +52,24 @@ namespace Style.setOption /-- Whether a syntax element is a `set_option` command, tactic or term: Return the name of the option being set, if any. -/ -def parse_set_option : Syntax → Option Name +def parseSetOption : Syntax → Option Name -- This handles all four possibilities of `_val`: a string, number, `true` and `false`. | `(command|set_option $name:ident $_val) => some name.getId | `(set_option $name:ident $_val in $_x) => some name.getId | `(tactic|set_option $name:ident $_val in $_x) => some name.getId | _ => none +/-- Deprecated alias for `Mathlib.Linter.Style.setOption.parseSetOption`. -/ +@[deprecated parseSetOption (since := "2024-12-07")] +def parse_set_option := @parseSetOption + /-- Whether a given piece of syntax is a `set_option` command, tactic or term. -/ -def is_set_option : Syntax → Bool := - fun stx ↦ parse_set_option stx matches some _name +def isSetOption : Syntax → Bool := + fun stx ↦ parseSetOption stx matches some _name + +/-- Deprecated alias for `Mathlib.Linter.Style.setOption.isSetOption`. -/ +@[deprecated isSetOption (since := "2024-12-07")] +def is_set_option := @isSetOption /-- The `setOption` linter: this lints any `set_option` command, term or tactic which sets a `pp`, `profiler` or `trace` option. @@ -76,8 +84,8 @@ def setOptionLinter : Linter where run := withSetOptionIn fun stx => do return if (← MonadState.get).messages.hasErrors then return - if let some head := stx.find? is_set_option then - if let some name := parse_set_option head then + if let some head := stx.find? isSetOption then + if let some name := parseSetOption head then let forbidden := [`debug, `pp, `profiler, `trace] if forbidden.contains name.getRoot then Linter.logLint linter.style.setOption head diff --git a/Mathlib/Tactic/MoveAdd.lean b/Mathlib/Tactic/MoveAdd.lean index 12f7051a82e2c..76c4709fd49a4 100644 --- a/Mathlib/Tactic/MoveAdd.lean +++ b/Mathlib/Tactic/MoveAdd.lean @@ -336,7 +336,7 @@ def pairUp : List (Expr × Bool × Syntax) → List Expr → To support a new binary operation, extend the list in this definition, so that it contains enough lemmas to allow `simp` to close a generic permutation goal for the new binary operation. -/ -def move_oper_simpCtx : MetaM Simp.Context := do +def moveOperSimpCtx : MetaM Simp.Context := do let simpNames := Elab.Tactic.simpOnlyBuiltins ++ [ ``add_comm, ``add_assoc, ``add_left_comm, -- for `HAdd.hAdd` ``mul_comm, ``mul_assoc, ``mul_left_comm, -- for `HMul.hMul` @@ -348,6 +348,9 @@ def move_oper_simpCtx : MetaM Simp.Context := do let simpThms ← simpNames.foldlM (·.addConst ·) ({} : SimpTheorems) Simp.mkContext {} (simpTheorems := #[simpThms]) +@[deprecated (since := "2024-12-07")] +alias move_oper_simpCtx := moveOperSimpCtx + /-- `reorderAndSimp mv op instr` takes as input an `MVarId` `mv`, the name `op` of a binary operation and a list of "instructions" `instr` that it passes to `permuteExpr`. @@ -368,7 +371,7 @@ def reorderAndSimp (mv : MVarId) (instr : List (Expr × Bool)) : throwError m!"There should only be 2 goals, instead of {twoGoals.length}" -- `permGoal` is the single goal `mv_permuted`, possibly more operations will be permuted later on let permGoal ← twoGoals.filterM fun v => return !(← v.isAssigned) - match ← (simpGoal (permGoal[1]!) (← move_oper_simpCtx)) with + match ← (simpGoal (permGoal[1]!) (← moveOperSimpCtx)) with | (some x, _) => throwError m!"'move_oper' could not solve {indentD x.2}" | (none, _) => return permGoal From 65c00859d05eb8ff717bb0f06ff355609be271b2 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Sun, 8 Dec 2024 17:47:11 +0000 Subject: [PATCH 573/829] chore: update Mathlib dependencies 2024-12-08 (#19813) 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 31395f1077639..df098d92919ba 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "c016aa9938c4cedc9b7066099f99bcae1b1af625", + "rev": "74dffd1a83cdd2969a31c9892b0517e7c6f50668", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From bd9aea4b2403dcb1ec3998335f71109212ada88e Mon Sep 17 00:00:00 2001 From: FMLJohn <417963616@qq.com> Date: Mon, 9 Dec 2024 08:03:26 +0000 Subject: [PATCH 574/829] feat(RingTheory/Polynomial/HilbertPoly): `Polynomial.natDegree_hilbertPoly_of_ne_zero` (#19765) --- Mathlib/Algebra/Polynomial/Degree/Domain.lean | 9 ++ .../RingTheory/Polynomial/HilbertPoly.lean | 115 ++++++++++++++---- 2 files changed, 101 insertions(+), 23 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/Degree/Domain.lean b/Mathlib/Algebra/Polynomial/Degree/Domain.lean index cad96a6912720..04519630e17d6 100644 --- a/Mathlib/Algebra/Polynomial/Degree/Domain.lean +++ b/Mathlib/Algebra/Polynomial/Degree/Domain.lean @@ -40,6 +40,15 @@ lemma natDegree_mul (hp : p ≠ 0) (hq : q ≠ 0) : (p*q).natDegree = p.natDegre rw [← Nat.cast_inj (R := WithBot ℕ), ← degree_eq_natDegree (mul_ne_zero hp hq), Nat.cast_add, ← degree_eq_natDegree hp, ← degree_eq_natDegree hq, degree_mul] +variable (p) in +lemma natDegree_smul (ha : a ≠ 0) : (a • p).natDegree = p.natDegree := by + by_cases hp : p = 0 + · simp only [hp, smul_zero] + · apply natDegree_eq_of_le_of_coeff_ne_zero + · exact (natDegree_smul_le _ _).trans (le_refl _) + · simpa only [coeff_smul, coeff_natDegree, smul_eq_mul, ne_eq, mul_eq_zero, + leadingCoeff_eq_zero, not_or] using ⟨ha, hp⟩ + @[simp] lemma natDegree_pow (p : R[X]) (n : ℕ) : natDegree (p ^ n) = n * natDegree p := by classical diff --git a/Mathlib/RingTheory/Polynomial/HilbertPoly.lean b/Mathlib/RingTheory/Polynomial/HilbertPoly.lean index 7aa04415b3b93..806c1a591b0d2 100644 --- a/Mathlib/RingTheory/Polynomial/HilbertPoly.lean +++ b/Mathlib/RingTheory/Polynomial/HilbertPoly.lean @@ -58,9 +58,33 @@ equals `(n - k + d).choose d`. noncomputable def preHilbertPoly (d k : ℕ) : F[X] := (d.factorial : F)⁻¹ • ((ascPochhammer F d).comp (Polynomial.X - (C (k : F)) + 1)) +lemma natDegree_preHilbertPoly [CharZero F] (d k : ℕ) : + (preHilbertPoly F d k).natDegree = d := by + have hne : (d ! : F) ≠ 0 := by norm_cast; positivity + rw [preHilbertPoly, natDegree_smul _ (inv_ne_zero hne), natDegree_comp, ascPochhammer_natDegree, + add_comm_sub, ← C_1, ← map_sub, natDegree_add_C, natDegree_X, mul_one] + +lemma coeff_preHilbertPoly_self [CharZero F] (d k : ℕ) : + (preHilbertPoly F d k).coeff d = (d ! : F)⁻¹ := by + delta preHilbertPoly + have hne : (d ! : F) ≠ 0 := by norm_cast; positivity + have heq : d = ((ascPochhammer F d).comp (X - C (k : F) + 1)).natDegree := + (natDegree_preHilbertPoly F d k).symm.trans (natDegree_smul _ (inv_ne_zero hne)) + nth_rw 3 [heq] + calc + _ = (d ! : F)⁻¹ • ((ascPochhammer F d).comp (X - C ((k : F) - 1))).leadingCoeff := by + simp only [sub_add, ← C_1, ← map_sub, coeff_smul, coeff_natDegree] + _ = (d ! : F)⁻¹ := by + simp only [leadingCoeff_comp (ne_of_eq_of_ne (natDegree_X_sub_C _) one_ne_zero), Monic.def.1 + (monic_ascPochhammer _ _), one_mul, leadingCoeff_X_sub_C, one_pow, smul_eq_mul, mul_one] + +lemma leadingCoeff_preHilbertPoly [CharZero F] (d k : ℕ) : + (preHilbertPoly F d k).leadingCoeff = (d ! : F)⁻¹ := by + rw [leadingCoeff, natDegree_preHilbertPoly, coeff_preHilbertPoly_self] + lemma preHilbertPoly_eq_choose_sub_add [CharZero F] (d : ℕ) {k n : ℕ} (hkn : k ≤ n): (preHilbertPoly F d k).eval (n : F) = (n - k + d).choose d := by - have : ((d ! : ℕ) : F) ≠ 0 := by norm_cast; positivity + have : (d ! : F) ≠ 0 := by norm_cast; positivity calc _ = (↑d !)⁻¹ * eval (↑(n - k + 1)) (ascPochhammer F d) := by simp [cast_sub hkn, preHilbertPoly] _ = (n - k + d).choose d := by @@ -75,7 +99,7 @@ is defined as `∑ i in p.support, (p.coeff i) • Polynomial.preHilbertPoly F d `M` is a graded module whose Poincaré series can be written as `p(X)/(1 - X)ᵈ` for some `p : ℚ[X]` with integer coefficients, then `Polynomial.hilbertPoly p d` is the Hilbert polynomial of `M`. See also `Polynomial.coeff_mul_invOneSubPow_eq_hilbertPoly_eval`, -which says that `PowerSeries.coeff F n (p * (PowerSeries.invOneSubPow F d))` equals +which says that `PowerSeries.coeff F n (p * PowerSeries.invOneSubPow F d)` equals `(Polynomial.hilbertPoly p d).eval (n : F)` for any large enough `n : ℕ`. -/ noncomputable def hilbertPoly (p : F[X]) : (d : ℕ) → F[X] @@ -107,7 +131,7 @@ coefficient of `Xⁿ` in the power series expansion of `p/(1 - X)ᵈ`. -/ theorem coeff_mul_invOneSubPow_eq_hilbertPoly_eval {p : F[X]} (d : ℕ) {n : ℕ} (hn : p.natDegree < n) : - PowerSeries.coeff F n (p * (invOneSubPow F d)) = (hilbertPoly p d).eval (n : F) := by + PowerSeries.coeff F n (p * invOneSubPow F d) = (hilbertPoly p d).eval (n : F) := by delta hilbertPoly; induction d with | zero => simp only [invOneSubPow_zero, Units.val_one, mul_one, coeff_coe, eval_zero] exact coeff_eq_zero_of_natDegree_lt hn @@ -132,14 +156,11 @@ theorem coeff_mul_invOneSubPow_eq_hilbertPoly_eval · rw [add_comm, Nat.add_sub_assoc (h_le ⟨x, hx⟩), succ_eq_add_one, add_tsub_cancel_right] /-- -The polynomial satisfying the key property of `Polynomial.hilbertPoly p d` is unique. In other -words, if `h : F[X]` and there exists some `N : ℕ` such that for any number `n : ℕ` bigger than -`N` we have `PowerSeries.coeff F n (p * (invOneSubPow F d)) = h.eval (n : F)`, then `h` is exactly -`Polynomial.hilbertPoly p d`. +The polynomial satisfying the key property of `Polynomial.hilbertPoly p d` is unique. -/ theorem exists_unique_hilbertPoly (p : F[X]) (d : ℕ) : - ∃! (h : F[X]), (∃ (N : ℕ), (∀ (n : ℕ) (_ : N < n), - PowerSeries.coeff F n (p * (invOneSubPow F d)) = h.eval (n : F))) := by + ∃! h : F[X], ∃ N : ℕ, ∀ n > N, + PowerSeries.coeff F n (p * invOneSubPow F d) = h.eval (n : F) := by use hilbertPoly p d; constructor · use p.natDegree exact fun n => coeff_mul_invOneSubPow_eq_hilbertPoly_eval d @@ -150,22 +171,25 @@ theorem exists_unique_hilbertPoly (p : F[X]) (d : ℕ) : simp only [Set.mem_Ioi, sup_lt_iff, Set.mem_setOf_eq] at hn ⊢ rw [← coeff_mul_invOneSubPow_eq_hilbertPoly_eval d hn.2, hhN n hn.1] +/-- +If `h : F[X]` and there exists some `N : ℕ` such that for any number `n : ℕ` bigger than `N` +we have `PowerSeries.coeff F n (p * invOneSubPow F d) = h.eval (n : F)`, then `h` is exactly +`Polynomial.hilbertPoly p d`. +-/ +theorem eq_hilbertPoly_of_forall_coeff_eq_eval + {p h : F[X]} {d : ℕ} (N : ℕ) (hhN : ∀ n > N, + PowerSeries.coeff F n (p * invOneSubPow F d) = h.eval (n : F)) : + h = hilbertPoly p d := + ExistsUnique.unique (exists_unique_hilbertPoly p d) ⟨N, hhN⟩ + ⟨p.natDegree, fun _ x => coeff_mul_invOneSubPow_eq_hilbertPoly_eval d x⟩ + lemma hilbertPoly_mul_one_sub_succ (p : F[X]) (d : ℕ) : hilbertPoly (p * (1 - X)) (d + 1) = hilbertPoly p d := by - apply eq_of_infinite_eval_eq - apply ((Set.Ioi_infinite (p * (1 - X)).natDegree).image cast_injective.injOn).mono - rintro x ⟨n, hn, rfl⟩ - simp only [Set.mem_setOf_eq] - by_cases hp : p = 0 - · simp only [hp, zero_mul, hilbertPoly_zero_nat] - · simp only [Set.mem_Ioi] at hn - have hpn : p.natDegree < n := by - suffices (1 : F[X]) - X ≠ 0 from lt_of_add_right_lt (natDegree_mul hp this ▸ hn) - rw [sub_ne_zero, ne_eq, ext_iff, not_forall] - use 0 - simp - simp only [hn, ← coeff_mul_invOneSubPow_eq_hilbertPoly_eval, coe_mul, coe_sub, coe_one, coe_X, - mul_assoc, hpn, ← one_sub_pow_mul_invOneSubPow_val_add_eq_invOneSubPow_val F d 1, pow_one] + apply eq_hilbertPoly_of_forall_coeff_eq_eval (p * (1 - X)).natDegree + intro n hn + have heq : 1 - PowerSeries.X = ((1 - X : F[X]) : F⟦X⟧) := by simp only [coe_sub, coe_one, coe_X] + rw [← one_sub_pow_mul_invOneSubPow_val_add_eq_invOneSubPow_val F d 1, pow_one, ← mul_assoc, heq, + ← coe_mul, coeff_mul_invOneSubPow_eq_hilbertPoly_eval (d + 1) hn] lemma hilbertPoly_mul_one_sub_pow_add (p : F[X]) (d e : ℕ) : hilbertPoly (p * (1 - X) ^ e) (d + e) = hilbertPoly p d := by @@ -173,4 +197,49 @@ lemma hilbertPoly_mul_one_sub_pow_add (p : F[X]) (d e : ℕ) : | zero => simp | succ e he => rw [pow_add, pow_one, ← mul_assoc, ← add_assoc, hilbertPoly_mul_one_sub_succ, he] +lemma hilbertPoly_eq_zero_of_le_rootMultiplicity_one + {p : F[X]} {d : ℕ} (hdp : d ≤ p.rootMultiplicity 1) : + hilbertPoly p d = 0 := by + by_cases hp : p = 0 + · rw [hp, hilbertPoly_zero_nat] + · rcases exists_eq_pow_rootMultiplicity_mul_and_not_dvd p hp 1 with ⟨q, hq1, hq2⟩ + have heq : p = q * (- 1) ^ p.rootMultiplicity 1 * (1 - X) ^ p.rootMultiplicity 1 := by + simp only [mul_assoc, ← mul_pow, neg_mul, one_mul, neg_sub] + exact hq1.trans (mul_comm _ _) + rw [heq, ← zero_add d, ← Nat.sub_add_cancel hdp, pow_add (1 - X), ← mul_assoc, + hilbertPoly_mul_one_sub_pow_add, hilbertPoly] + +theorem natDegree_hilbertPoly_of_ne_zero_of_rootMultiplicity_lt + {p : F[X]} {d : ℕ} (hp : p ≠ 0) (hpd : p.rootMultiplicity 1 < d) : + (hilbertPoly p d).natDegree = d - p.rootMultiplicity 1 - 1 := by + rcases exists_eq_pow_rootMultiplicity_mul_and_not_dvd p hp 1 with ⟨q, hq1, hq2⟩ + have heq : p = q * (- 1) ^ p.rootMultiplicity 1 * (1 - X) ^ p.rootMultiplicity 1 := by + simp only [mul_assoc, ← mul_pow, neg_mul, one_mul, neg_sub] + exact hq1.trans (mul_comm _ _) + nth_rw 1 [heq, ← Nat.sub_add_cancel (le_of_lt hpd), hilbertPoly_mul_one_sub_pow_add, + ← Nat.sub_add_cancel (Nat.le_sub_of_add_le' <| add_one_le_of_lt hpd)] + delta hilbertPoly + apply natDegree_eq_of_le_of_coeff_ne_zero + · apply natDegree_sum_le_of_forall_le _ _ <| fun _ _ => ?_ + apply le_trans (natDegree_smul_le _ _) + rw [natDegree_preHilbertPoly] + · have : (fun (x : ℕ) (a : F) => a) = fun x a => a * 1 ^ x := by simp only [one_pow, mul_one] + simp only [finset_sum_coeff, coeff_smul, smul_eq_mul, coeff_preHilbertPoly_self, + ← Finset.sum_mul, ← sum_def _ (fun _ a => a), this, ← eval_eq_sum, eval_mul, eval_pow, + eval_neg, eval_one, _root_.mul_eq_zero, pow_eq_zero_iff', neg_eq_zero, one_ne_zero, ne_eq, + false_and, or_false, inv_eq_zero, cast_eq_zero, not_or] + exact ⟨(not_iff_not.2 dvd_iff_isRoot).1 hq2, factorial_ne_zero _⟩ + +theorem natDegree_hilbertPoly_of_ne_zero + {p : F[X]} {d : ℕ} (hh : hilbertPoly p d ≠ 0) : + (hilbertPoly p d).natDegree = d - p.rootMultiplicity 1 - 1 := by + have hp : p ≠ 0 := by + intro h + rw [h] at hh + exact hh (hilbertPoly_zero_nat F d) + have hpd : p.rootMultiplicity 1 < d := by + by_contra h + exact hh (hilbertPoly_eq_zero_of_le_rootMultiplicity_one <| not_lt.1 h) + exact natDegree_hilbertPoly_of_ne_zero_of_rootMultiplicity_lt hp hpd + end Polynomial From d592ae383dff29cdfaa565eda6dc561c69e69ec7 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Mon, 9 Dec 2024 08:25:42 +0000 Subject: [PATCH 575/829] feat(LinearAlgebra/Goursat): Goursat's lemma on submodules of products (#18667) This proves Goursat's lemma for modules, characterising submodules of `A x B` as graphs of isomorphisms between subquotients of `A` and `B`. --- Mathlib.lean | 1 + Mathlib/LinearAlgebra/Goursat.lean | 137 +++++++++++++++++++++++++++++ Mathlib/LinearAlgebra/Prod.lean | 4 + 3 files changed, 142 insertions(+) create mode 100644 Mathlib/LinearAlgebra/Goursat.lean diff --git a/Mathlib.lean b/Mathlib.lean index e1579076cc334..3aa01e690df3d 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3362,6 +3362,7 @@ import Mathlib.LinearAlgebra.FreeModule.PID import Mathlib.LinearAlgebra.FreeModule.StrongRankCondition import Mathlib.LinearAlgebra.FreeProduct.Basic import Mathlib.LinearAlgebra.GeneralLinearGroup +import Mathlib.LinearAlgebra.Goursat import Mathlib.LinearAlgebra.InvariantBasisNumber import Mathlib.LinearAlgebra.Isomorphisms import Mathlib.LinearAlgebra.JordanChevalley diff --git a/Mathlib/LinearAlgebra/Goursat.lean b/Mathlib/LinearAlgebra/Goursat.lean new file mode 100644 index 0000000000000..0316c6952d5d5 --- /dev/null +++ b/Mathlib/LinearAlgebra/Goursat.lean @@ -0,0 +1,137 @@ +/- +Copyright (c) 2024 David Loeffler. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Loeffler +-/ +import Mathlib.GroupTheory.Goursat +import Mathlib.LinearAlgebra.Prod +import Mathlib.LinearAlgebra.Quotient.Basic + +/-! +# Goursat's lemma for submodules + +Let `M, N` be modules over a ring `R`. If `L` is a submodule of `M × N` which projects fully on +to both factors, then there exist submodules `M' ≤ M` and `N' ≤ N` such that `M' × N' ≤ L` and the +image of `L` in `(M ⧸ M') × (N ⧸ N')` is the graph of an isomorphism `M ⧸ M' ≃ₗ[R] N ⧸ N'`. +Equivalently, `L` is equal to the preimage in `M × N` of the graph of this isomorphism +`M ⧸ M' ≃ₗ[R] N ⧸ N'`. + +`M'` and `N'` can be explicitly constructed as `Submodule.goursatFst L` and `Submodule.goursatSnd L` +respectively. +-/ + +open Function Set LinearMap + +namespace Submodule +variable {R M N : Type*} [Ring R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] + {L : Submodule R (M × N)} + (hL₁ : Surjective (Prod.fst ∘ L.subtype)) (hL₂ : Surjective (Prod.snd ∘ L.subtype)) + +variable (L) in +/-- For `L` a submodule of `M × N`, `L.goursatFst` is the kernel of the projection map `L → N`, +considered as a submodule of `M`. + +This is the first submodule appearing in Goursat's lemma. See `Subgroup.goursat`. -/ +def goursatFst : Submodule R M := + ((LinearMap.snd R M N).comp L.subtype).ker.map ((LinearMap.fst R M N).comp L.subtype) + + +variable (L) in +/-- For `L` a subgroup of `M × N`, `L.goursatSnd` is the kernel of the projection map `L → M`, +considered as a subgroup of `N`. + +This is the second subgroup appearing in Goursat's lemma. See `Subgroup.goursat`. -/ +def goursatSnd : Submodule R N := + ((LinearMap.fst R M N).comp L.subtype).ker.map ((LinearMap.snd R M N).comp L.subtype) + +lemma goursatFst_toAddSubgroup : + (goursatFst L).toAddSubgroup = L.toAddSubgroup.goursatFst := by + ext x + simp [mem_toAddSubgroup, goursatFst, AddSubgroup.mem_goursatFst] + +lemma goursatSnd_toAddSubgroup : + (goursatSnd L).toAddSubgroup = L.toAddSubgroup.goursatSnd := by + ext x + simp [mem_toAddSubgroup, goursatSnd, AddSubgroup.mem_goursatSnd] + +variable (L) in +lemma goursatFst_prod_goursatSnd_le : L.goursatFst.prod L.goursatSnd ≤ L := by + simpa only [← toAddSubgroup_le, goursatFst_toAddSubgroup, goursatSnd_toAddSubgroup] + using L.toAddSubgroup.goursatFst_prod_goursatSnd_le + +include hL₁ hL₂ in +/-- **Goursat's lemma** for a submodule of a product with surjective projections. + +If `L` is a submodule of `M × N` which projects fully on both factors, then there exist submodules +`M' ≤ M` and `N' ≤ N` such that `M' × N' ≤ L` and the image of `L` in `(M ⧸ M') × (N ⧸ N')` is the +graph of an isomorphism of `R`-modules `(M ⧸ M') ≃ (N ⧸ N')`. + +`M` and `N` can be explicitly constructed as `L.goursatFst` and `L.goursatSnd` respectively. -/ +lemma goursat_surjective : ∃ e : (M ⧸ L.goursatFst) ≃ₗ[R] N ⧸ L.goursatSnd, + LinearMap.range ((L.goursatFst.mkQ.prodMap L.goursatSnd.mkQ).comp L.subtype) = e.graph := by + -- apply add-group result + obtain ⟨(e : M ⧸ L.goursatFst ≃+ N ⧸ L.goursatSnd), he⟩ := + L.toAddSubgroup.goursat_surjective hL₁ hL₂ + -- check R-linearity of the map + have (r : R) (x : M ⧸ L.goursatFst) : e (r • x) = r • e x := by + show (r • x, r • e x) ∈ e.toAddMonoidHom.graph + rw [← he, ← Prod.smul_mk] + have : (x, e x) ∈ e.toAddMonoidHom.graph := rfl + rw [← he, AddMonoidHom.mem_range] at this + rcases this with ⟨⟨l, hl⟩, hl'⟩ + use ⟨r • l, L.smul_mem r hl⟩ + rw [← hl'] + rfl + -- define the map as an R-linear equiv + use { e with map_smul' := this } + rw [← toAddSubgroup_injective.eq_iff] + convert he using 1 + ext v + rw [mem_toAddSubgroup, mem_graph_iff, Eq.comm] + rfl + +/-- **Goursat's lemma** for an arbitrary submodule of a product. + +If `L` is a submodule of `M × N`, then there exist submodules `M'' ≤ M' ≤ M` and `N'' ≤ N' ≤ N` such +that `L ≤ M' × N'`, and `L` is (the image in `M × N` of) the preimage of the graph of an `R`-linear +isomorphism `M' ⧸ M'' ≃ N' ⧸ N''`. -/ +lemma goursat : ∃ (M' : Submodule R M) (N' : Submodule R N) (M'' : Submodule R M') + (N'' : Submodule R N') (e : (M' ⧸ M'') ≃ₗ[R] N' ⧸ N''), + L = (e.graph.comap <| M''.mkQ.prodMap N''.mkQ).map (M'.subtype.prodMap N'.subtype) := by + let M' := L.map (LinearMap.fst ..) + let N' := L.map (LinearMap.snd ..) + let P : L →ₗ[R] M' := (LinearMap.fst ..).submoduleMap L + let Q : L →ₗ[R] N' := (LinearMap.snd ..).submoduleMap L + let L' : Submodule R (M' × N') := LinearMap.range (P.prod Q) + have hL₁' : Surjective (Prod.fst ∘ L'.subtype) := by + simp only [← coe_fst (R := R), ← coe_comp, ← range_eq_top, LinearMap.range_comp, range_subtype] + simpa only [L', ← LinearMap.range_comp, fst_prod, range_eq_top] using + (LinearMap.fst ..).submoduleMap_surjective L + have hL₂' : Surjective (Prod.snd ∘ L'.subtype) := by + simp only [← coe_snd (R := R), ← coe_comp, ← range_eq_top, LinearMap.range_comp, range_subtype] + simpa only [L', ← LinearMap.range_comp, snd_prod, range_eq_top] using + (LinearMap.snd ..).submoduleMap_surjective L + obtain ⟨e, he⟩ := goursat_surjective hL₁' hL₂' + use M', N', L'.goursatFst, L'.goursatSnd, e + rw [← he] + simp only [LinearMap.range_comp, Submodule.range_subtype, L'] + rw [comap_map_eq_self] + · ext ⟨m, n⟩ + constructor + · intro hmn + simp only [mem_map, LinearMap.mem_range, prod_apply, Subtype.exists, Prod.exists, coe_prodMap, + coe_subtype, Prod.map_apply, Prod.mk.injEq, exists_and_right, exists_eq_right_right, + exists_eq_right, M', N', fst_apply, snd_apply] + exact ⟨⟨n, hmn⟩, ⟨m, hmn⟩, ⟨m, n, hmn, rfl⟩⟩ + · simp only [mem_map, LinearMap.mem_range, prod_apply, Subtype.exists, Prod.exists, + coe_prodMap, coe_subtype, Prod.map_apply, Prod.mk.injEq, exists_and_right, + exists_eq_right_right, exists_eq_right, forall_exists_index, Pi.prod] + rintro hm hn m₁ n₁ hm₁n₁ ⟨hP, hQ⟩ + simp only [Subtype.ext_iff] at hP hQ + rwa [← hP, ← hQ] + · convert goursatFst_prod_goursatSnd_le (range <| P.prod Q) + ext ⟨m, n⟩ + simp_rw [mem_ker, coe_prodMap, Prod.map_apply, Submodule.mem_prod, Prod.zero_eq_mk, + Prod.ext_iff, ← mem_ker, ker_mkQ] + +end Submodule diff --git a/Mathlib/LinearAlgebra/Prod.lean b/Mathlib/LinearAlgebra/Prod.lean index bae77e00fdd50..8aa773428cf39 100644 --- a/Mathlib/LinearAlgebra/Prod.lean +++ b/Mathlib/LinearAlgebra/Prod.lean @@ -76,6 +76,10 @@ theorem fst_apply (x : M × M₂) : fst R M M₂ x = x.1 := theorem snd_apply (x : M × M₂) : snd R M M₂ x = x.2 := rfl +@[simp, norm_cast] lemma coe_fst : ⇑(fst R M M₂) = Prod.fst := rfl + +@[simp, norm_cast] lemma coe_snd : ⇑(snd R M M₂) = Prod.snd := rfl + theorem fst_surjective : Function.Surjective (fst R M M₂) := fun x => ⟨(x, 0), rfl⟩ theorem snd_surjective : Function.Surjective (snd R M M₂) := fun x => ⟨(0, x), rfl⟩ From bde99e32878343e11a019831a0fda0fa5490fa51 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 9 Dec 2024 08:25:43 +0000 Subject: [PATCH 576/829] chore: don't import `Set.indicator` in `Topology.ContinuousOn` (#19653) This is a first step towards not depending on algebra in basic topology --- Mathlib.lean | 1 + Mathlib/Topology/Algebra/Indicator.lean | 38 +++++++++++++++++++ .../Topology/ContinuousMap/Bounded/Basic.lean | 5 ++- Mathlib/Topology/ContinuousOn.lean | 28 -------------- Mathlib/Topology/IndicatorConstPointwise.lean | 1 + Mathlib/Topology/Order/LeftRight.lean | 1 + Mathlib/Topology/Separation/Basic.lean | 1 + 7 files changed, 45 insertions(+), 30 deletions(-) create mode 100644 Mathlib/Topology/Algebra/Indicator.lean diff --git a/Mathlib.lean b/Mathlib.lean index 3aa01e690df3d..69e93b3a82b6e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4961,6 +4961,7 @@ import Mathlib.Topology.Algebra.Group.SubmonoidClosure import Mathlib.Topology.Algebra.Group.TopologicalAbelianization import Mathlib.Topology.Algebra.GroupCompletion import Mathlib.Topology.Algebra.GroupWithZero +import Mathlib.Topology.Algebra.Indicator import Mathlib.Topology.Algebra.InfiniteSum.Basic import Mathlib.Topology.Algebra.InfiniteSum.Constructions import Mathlib.Topology.Algebra.InfiniteSum.Defs diff --git a/Mathlib/Topology/Algebra/Indicator.lean b/Mathlib/Topology/Algebra/Indicator.lean new file mode 100644 index 0000000000000..189c9d6f29038 --- /dev/null +++ b/Mathlib/Topology/Algebra/Indicator.lean @@ -0,0 +1,38 @@ +/- +Copyright (c) 2024 PFR contributors. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: PFR contributors +-/ +import Mathlib.Algebra.Group.Indicator +import Mathlib.Topology.ContinuousOn + +/-! +# Continuity of indicator functions +-/ + +open Set +open scoped Topology + +variable {α β : Type*} [TopologicalSpace α] [TopologicalSpace β] {f : α → β} {s : Set α} [One β] + +@[to_additive] +lemma continuous_mulIndicator (hs : ∀ a ∈ frontier s, f a = 1) (hf : ContinuousOn f (closure s)) : + Continuous (mulIndicator s f) := by + classical exact continuous_piecewise hs hf continuousOn_const + +@[to_additive] +protected lemma Continuous.mulIndicator (hs : ∀ a ∈ frontier s, f a = 1) (hf : Continuous f) : + Continuous (mulIndicator s f) := by + classical exact hf.piecewise hs continuous_const + +@[to_additive] +theorem ContinuousOn.continuousAt_mulIndicator (hf : ContinuousOn f (interior s)) {x : α} + (hx : x ∉ frontier s) : + ContinuousAt (s.mulIndicator f) x := by + rw [← Set.mem_compl_iff, compl_frontier_eq_union_interior] at hx + obtain h | h := hx + · have hs : interior s ∈ 𝓝 x := mem_interior_iff_mem_nhds.mp (by rwa [interior_interior]) + exact ContinuousAt.congr (hf.continuousAt hs) <| Filter.eventuallyEq_iff_exists_mem.mpr + ⟨interior s, hs, Set.eqOn_mulIndicator.symm.mono interior_subset⟩ + · exact ContinuousAt.congr continuousAt_const <| Filter.eventuallyEq_iff_exists_mem.mpr + ⟨sᶜ, mem_interior_iff_mem_nhds.mp h, Set.eqOn_mulIndicator'.symm⟩ diff --git a/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean b/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean index 81b40ea72345a..5c2ccedeed4ea 100644 --- a/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean +++ b/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean @@ -4,11 +4,12 @@ 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.ContinuousMap.Algebra import Mathlib.Analysis.Normed.Order.Lattice import Mathlib.Analysis.NormedSpace.OperatorNorm.Basic -import Mathlib.Topology.Bornology.BoundedOperation import Mathlib.Tactic.Monotonicity +import Mathlib.Topology.Algebra.Indicator +import Mathlib.Topology.Bornology.BoundedOperation +import Mathlib.Topology.ContinuousMap.Algebra /-! # Bounded continuous functions diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index d65b86c859fa0..56c76196a7d53 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -3,7 +3,6 @@ Copyright (c) 2019 Reid Barton. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import Mathlib.Algebra.Group.Indicator import Mathlib.Topology.Constructions /-! @@ -1403,33 +1402,6 @@ theorem Continuous.piecewise [∀ a, Decidable (a ∈ s)] Continuous (piecewise s f g) := hf.if hs hg -section Indicator -variable [One β] - -@[to_additive] -lemma continuous_mulIndicator (hs : ∀ a ∈ frontier s, f a = 1) (hf : ContinuousOn f (closure s)) : - Continuous (mulIndicator s f) := by - classical exact continuous_piecewise hs hf continuousOn_const - -@[to_additive] -protected lemma Continuous.mulIndicator (hs : ∀ a ∈ frontier s, f a = 1) (hf : Continuous f) : - Continuous (mulIndicator s f) := by - classical exact hf.piecewise hs continuous_const - -@[to_additive] -theorem ContinuousOn.continuousAt_mulIndicator (hf : ContinuousOn f (interior s)) {x : α} - (hx : x ∉ frontier s) : - ContinuousAt (s.mulIndicator f) x := by - rw [← Set.mem_compl_iff, compl_frontier_eq_union_interior] at hx - obtain h | h := hx - · have hs : interior s ∈ 𝓝 x := mem_interior_iff_mem_nhds.mp (by rwa [interior_interior]) - exact ContinuousAt.congr (hf.continuousAt hs) <| Filter.eventuallyEq_iff_exists_mem.mpr - ⟨interior s, hs, Set.eqOn_mulIndicator.symm.mono interior_subset⟩ - · exact ContinuousAt.congr continuousAt_const <| Filter.eventuallyEq_iff_exists_mem.mpr - ⟨sᶜ, mem_interior_iff_mem_nhds.mp h, Set.eqOn_mulIndicator'.symm⟩ - -end Indicator - theorem IsOpen.ite' (hs : IsOpen s) (hs' : IsOpen s') (ht : ∀ x ∈ frontier t, x ∈ s ↔ x ∈ s') : IsOpen (t.ite s s') := by classical diff --git a/Mathlib/Topology/IndicatorConstPointwise.lean b/Mathlib/Topology/IndicatorConstPointwise.lean index deb980b52a7e8..0ec0a97c72bba 100644 --- a/Mathlib/Topology/IndicatorConstPointwise.lean +++ b/Mathlib/Topology/IndicatorConstPointwise.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Kalle Kytölä. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kalle Kytölä -/ +import Mathlib.Algebra.Group.Indicator import Mathlib.Topology.Separation.Basic /-! diff --git a/Mathlib/Topology/Order/LeftRight.lean b/Mathlib/Topology/Order/LeftRight.lean index e18f75c1f4fa6..065632d5f1161 100644 --- a/Mathlib/Topology/Order/LeftRight.lean +++ b/Mathlib/Topology/Order/LeftRight.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Anatole Dedecker. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anatole Dedecker -/ +import Mathlib.Order.Antichain import Mathlib.Topology.ContinuousOn /-! diff --git a/Mathlib/Topology/Separation/Basic.lean b/Mathlib/Topology/Separation/Basic.lean index b56a3a27aa9a9..827ce1e31923c 100644 --- a/Mathlib/Topology/Separation/Basic.lean +++ b/Mathlib/Topology/Separation/Basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ +import Mathlib.Algebra.Group.Support import Mathlib.Topology.Compactness.Lindelof import Mathlib.Topology.Compactness.SigmaCompact import Mathlib.Topology.Connected.TotallyDisconnected From 415fc064d791d4fb77a74a636db6cfba0f75f224 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Mon, 9 Dec 2024 08:25:45 +0000 Subject: [PATCH 577/829] feat: add optional message for `try_this` tactic, and use it for `ring` (#19763) This adds a message to the `ring` tactic when `ring1` fails to solve the goal. It informs the user of common pitfalls when using `ring`, and suggests `noncomm_ring`, `abel` and `module` as alternatives. This confusion occurs repeatedly on Zulip, e.g., [here](https://leanprover.zulipchat.com/#narrow/channel/113489-new-members/topic/why.20ring.20doesn't.20work.20on.20ring/near/486365363) and [here](https://leanprover.zulipchat.com/#narrow/channel/113488-general/topic/Behaviour.20of.20ring.20tactic/near/436204738) to name only two examples. --- Mathlib/Tactic/Ring/RingNF.lean | 24 ++++++++++++++++++++---- Mathlib/Tactic/TryThis.lean | 12 ++++++++---- 2 files changed, 28 insertions(+), 8 deletions(-) diff --git a/Mathlib/Tactic/Ring/RingNF.lean b/Mathlib/Tactic/Ring/RingNF.lean index 3d57d68676c46..095a8b819b8d4 100644 --- a/Mathlib/Tactic/Ring/RingNF.lean +++ b/Mathlib/Tactic/Ring/RingNF.lean @@ -259,9 +259,17 @@ example (x : ℕ) (h : x * 2 > 5): x + x > 5 := by ring; assumption -- suggests ``` -/ macro (name := ring) "ring" : tactic => - `(tactic| first | ring1 | try_this ring_nf) + `(tactic| first | ring1 | try_this ring_nf + "\n\nThe `ring` tactic failed to close the goal. Use `ring_nf` to obtain a normal form. + \nNote that `ring` works primarily in *commutative* rings. \ + If you have a noncommutative ring, abelian group or module, consider using \ + `noncomm_ring`, `abel` or `module` instead.") @[inherit_doc ring] macro "ring!" : tactic => - `(tactic| first | ring1! | try_this ring_nf!) + `(tactic| first | ring1! | try_this ring_nf! + "\n\nThe `ring!` tactic failed to close the goal. Use `ring_nf!` to obtain a normal form. + \nNote that `ring!` works primarily in *commutative* rings. \ + If you have a noncommutative ring, abelian group or module, consider using \ + `noncomm_ring`, `abel` or `module` instead.") /-- The tactic `ring` evaluates expressions in *commutative* (semi)rings. @@ -270,9 +278,17 @@ This is the conv tactic version, which rewrites a target which is a ring equalit See also the `ring` tactic. -/ macro (name := ringConv) "ring" : conv => - `(conv| first | discharge => ring1 | try_this ring_nf) + `(conv| first | discharge => ring1 | try_this ring_nf + "\n\nThe `ring` tactic failed to close the goal. Use `ring_nf` to obtain a normal form. + \nNote that `ring` works primarily in *commutative* rings. \ + If you have a noncommutative ring, abelian group or module, consider using \ + `noncomm_ring`, `abel` or `module` instead.") @[inherit_doc ringConv] macro "ring!" : conv => - `(conv| first | discharge => ring1! | try_this ring_nf!) + `(conv| first | discharge => ring1! | try_this ring_nf! + "\n\nThe `ring!` tactic failed to close the goal. Use `ring_nf!` to obtain a normal form. + \nNote that `ring!` works primarily in *commutative* rings. \ + If you have a noncommutative ring, abelian group or module, consider using \ + `noncomm_ring`, `abel` or `module` instead.") end RingNF diff --git a/Mathlib/Tactic/TryThis.lean b/Mathlib/Tactic/TryThis.lean index 3d4814c3d4087..f34ea59970cb9 100644 --- a/Mathlib/Tactic/TryThis.lean +++ b/Mathlib/Tactic/TryThis.lean @@ -18,13 +18,17 @@ namespace Mathlib.Tactic open Lean /-- Produces the text `Try this: ` with the given tactic, and then executes it. -/ -elab tk:"try_this" tac:tactic : tactic => do +elab tk:"try_this" tac:tactic info:(str)? : tactic => do Elab.Tactic.evalTactic tac - Meta.Tactic.TryThis.addSuggestion tk tac (origSpan? := ← getRef) + Meta.Tactic.TryThis.addSuggestion tk + { suggestion := tac, postInfo? := TSyntax.getString <$> info } + (origSpan? := ← getRef) /-- Produces the text `Try this: ` with the given conv tactic, and then executes it. -/ -elab tk:"try_this" tac:conv : conv => do +elab tk:"try_this" tac:conv info:(str)? : conv => do Elab.Tactic.evalTactic tac - Meta.Tactic.TryThis.addSuggestion tk tac (origSpan? := ← getRef) + Meta.Tactic.TryThis.addSuggestion tk + { suggestion := tac, postInfo? := TSyntax.getString <$> info } + (origSpan? := ← getRef) end Mathlib.Tactic From cb03cd930be98a8f0ebdef586f8e9f8af625fc1b Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Mon, 9 Dec 2024 08:25:47 +0000 Subject: [PATCH 578/829] chore(RingTheory): golf `iInf_localization_eq_bot` (#19785) using roughly the same pattern as [`Submodule.mem_of_localization_maximal`](https://github.com/leanprover-community/mathlib4/commit/a5d04c838085d6c3554d59278c04126b62314757#diff-25704369c7e20d4992fa2ef4a09fc372db1df67c0e7c504d4a99820bb1ec8f46R40-R50) --- Mathlib/RingTheory/MaximalSpectrum.lean | 32 ++++++------------------- 1 file changed, 7 insertions(+), 25 deletions(-) diff --git a/Mathlib/RingTheory/MaximalSpectrum.lean b/Mathlib/RingTheory/MaximalSpectrum.lean index b9b3ccfdbd9db..959eae0a6633e 100644 --- a/Mathlib/RingTheory/MaximalSpectrum.lean +++ b/Mathlib/RingTheory/MaximalSpectrum.lean @@ -3,7 +3,6 @@ Copyright (c) 2022 David Kurniadi Angdinata. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Kurniadi Angdinata -/ -import Mathlib.RingTheory.Ideal.Colon import Mathlib.RingTheory.Localization.AsSubring import Mathlib.RingTheory.PrimeSpectrum @@ -70,23 +69,12 @@ theorem iInf_localization_eq_bot : (⨅ v : MaximalSpectrum R, constructor · contrapose intro hrange hlocal - let denom : Ideal R := (Submodule.span R {1} : Submodule R K).colon (Submodule.span R {x}) - have hdenom : (1 : R) ∉ denom := by - intro hdenom - rcases Submodule.mem_span_singleton.mp - (Submodule.mem_colon.mp hdenom x <| Submodule.mem_span_singleton_self x) with ⟨y, hy⟩ - exact hrange ⟨y, by rw [← mul_one <| algebraMap R K y, ← Algebra.smul_def, hy, one_smul]⟩ - rcases denom.exists_le_maximal fun h => (h ▸ hdenom) Submodule.mem_top with ⟨max, hmax, hle⟩ + let denom : Ideal R := (1 : Submodule R K).comap (LinearMap.toSpanSingleton R K x) + have hdenom : (1 : R) ∉ denom := by simpa [denom] using hrange + rcases denom.exists_le_maximal (denom.ne_top_iff_one.mpr hdenom) with ⟨max, hmax, hle⟩ rcases hlocal ⟨max, hmax⟩ with ⟨n, d, hd, rfl⟩ - apply hd (hle <| Submodule.mem_colon.mpr fun _ hy => _) - intro _ hy - rcases Submodule.mem_span_singleton.mp hy with ⟨y, rfl⟩ - exact Submodule.mem_span_singleton.mpr ⟨y * n, by - rw [Algebra.smul_def, mul_one, map_mul, smul_comm, Algebra.smul_def, Algebra.smul_def, - mul_comm <| algebraMap R K d, - inv_mul_cancel_right₀ <| - (map_ne_zero_iff _ <| NoZeroSMulDivisors.algebraMap_injective R K).mpr fun h => - (h ▸ hd) max.zero_mem]⟩ + exact hd (hle ⟨n, by simp [denom, Algebra.smul_def, mul_left_comm, mul_inv_cancel₀ <| + (map_ne_zero_iff _ <| IsFractionRing.injective R K).mpr fun h ↦ hd (h ▸ max.zero_mem :)]⟩) · rintro ⟨y, rfl⟩ ⟨v, hv⟩ exact ⟨y, 1, v.ne_top_iff_one.mp hv.ne_top, by rw [map_one, inv_one, mul_one]⟩ @@ -101,13 +89,7 @@ variable [IsDomain R] (K : Type v) [Field K] [Algebra R K] [IsFractionRing R K] viewed as subalgebras of its field of fractions. -/ theorem iInf_localization_eq_bot : ⨅ v : PrimeSpectrum R, Localization.subalgebra.ofField K _ (v.asIdeal.primeCompl_le_nonZeroDivisors) = ⊥ := by - ext x - rw [Algebra.mem_iInf] - constructor - · rw [← MaximalSpectrum.iInf_localization_eq_bot, Algebra.mem_iInf] - exact fun hx ⟨v, hv⟩ => hx ⟨v, hv.isPrime⟩ - · rw [Algebra.mem_bot] - rintro ⟨y, rfl⟩ ⟨v, hv⟩ - exact ⟨y, 1, v.ne_top_iff_one.mp hv.ne_top, by rw [map_one, inv_one, mul_one]⟩ + refine bot_unique (.trans (fun _ ↦ ?_) (MaximalSpectrum.iInf_localization_eq_bot R K).le) + simpa only [Algebra.mem_iInf] using fun hx ⟨v, hv⟩ ↦ hx ⟨v, hv.isPrime⟩ end PrimeSpectrum From 8f0293e98e7b732f1ee71fc511e9a6b74bf59041 Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Mon, 9 Dec 2024 08:25:49 +0000 Subject: [PATCH 579/829] chore(RingTheory/Flat): remove `Algebra.Flat` and move `RingHom.Flat` (#19786) We remove `Algebra.Flat` as it is a special case of `Module.Flat` and make `RingHom.Flat` a regular definition to align with e.g. `RingHom.Finite`. Also adds some meta properties of flat ring homomorphisms. Co-authored-by: Christian Merten <136261474+chrisflav@users.noreply.github.com> --- Mathlib.lean | 2 +- .../Algebra/Category/Ring/Under/Limits.lean | 3 +- .../AdicCompletion/AsTensorProduct.lean | 6 +- Mathlib/RingTheory/Flat/Algebra.lean | 85 ------------------- Mathlib/RingTheory/RingHom/Flat.lean | 71 ++++++++++++++++ 5 files changed, 77 insertions(+), 90 deletions(-) delete mode 100644 Mathlib/RingTheory/Flat/Algebra.lean create mode 100644 Mathlib/RingTheory/RingHom/Flat.lean diff --git a/Mathlib.lean b/Mathlib.lean index 69e93b3a82b6e..4ec2f1daeaf33 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4315,7 +4315,6 @@ import Mathlib.RingTheory.Finiteness.Projective import Mathlib.RingTheory.Finiteness.Subalgebra import Mathlib.RingTheory.Finiteness.TensorProduct import Mathlib.RingTheory.Fintype -import Mathlib.RingTheory.Flat.Algebra import Mathlib.RingTheory.Flat.Basic import Mathlib.RingTheory.Flat.CategoryTheory import Mathlib.RingTheory.Flat.Equalizer @@ -4539,6 +4538,7 @@ import Mathlib.RingTheory.Regular.RegularSequence import Mathlib.RingTheory.RingHom.Finite import Mathlib.RingTheory.RingHom.FinitePresentation import Mathlib.RingTheory.RingHom.FiniteType +import Mathlib.RingTheory.RingHom.Flat import Mathlib.RingTheory.RingHom.Injective import Mathlib.RingTheory.RingHom.Integral import Mathlib.RingTheory.RingHom.Locally diff --git a/Mathlib/Algebra/Category/Ring/Under/Limits.lean b/Mathlib/Algebra/Category/Ring/Under/Limits.lean index 810501a4eca05..29babdcccee80 100644 --- a/Mathlib/Algebra/Category/Ring/Under/Limits.lean +++ b/Mathlib/Algebra/Category/Ring/Under/Limits.lean @@ -7,7 +7,7 @@ import Mathlib.Algebra.Category.Ring.Under.Basic import Mathlib.CategoryTheory.Limits.Constructions.LimitsOfProductsAndEqualizers import Mathlib.CategoryTheory.Limits.Over import Mathlib.RingTheory.TensorProduct.Pi -import Mathlib.RingTheory.Flat.Algebra +import Mathlib.RingTheory.RingHom.Flat import Mathlib.RingTheory.Flat.Equalizer /-! @@ -208,6 +208,7 @@ lemma preservesFiniteLimits_of_flat (hf : RingHom.Flat f) : PreservesFiniteLimits (Under.pushout f) where preservesFiniteLimits _ := letI : Algebra R S := RingHom.toAlgebra f + haveI : Module.Flat R S := hf preservesLimitsOfShape_of_natIso (tensorProdIsoPushout R S) end CommRingCat.Under diff --git a/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean b/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean index af49d74e4ff2a..7327b43e29424 100644 --- a/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean +++ b/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean @@ -8,7 +8,7 @@ import Mathlib.CategoryTheory.Abelian.DiagramLemmas.Four import Mathlib.LinearAlgebra.TensorProduct.Pi import Mathlib.LinearAlgebra.TensorProduct.RightExactness import Mathlib.RingTheory.AdicCompletion.Exactness -import Mathlib.RingTheory.Flat.Algebra +import Mathlib.RingTheory.Flat.Basic /-! @@ -375,8 +375,8 @@ lemma tensor_map_id_left_injective_of_injective (hf : Function.Injective f) : end /-- Adic completion of a Noetherian ring `R` is flat over `R`. -/ -instance flat_of_isNoetherian [IsNoetherianRing R] : Algebra.Flat R (AdicCompletion I R) where - out := (Module.Flat.iff_lTensor_injective' R (AdicCompletion I R)).mpr <| fun J ↦ +instance flat_of_isNoetherian [IsNoetherianRing R] : Module.Flat R (AdicCompletion I R) := + (Module.Flat.iff_lTensor_injective' R (AdicCompletion I R)).mpr fun J ↦ tensor_map_id_left_injective_of_injective I (Submodule.injective_subtype J) end Noetherian diff --git a/Mathlib/RingTheory/Flat/Algebra.lean b/Mathlib/RingTheory/Flat/Algebra.lean deleted file mode 100644 index bf379a8806436..0000000000000 --- a/Mathlib/RingTheory/Flat/Algebra.lean +++ /dev/null @@ -1,85 +0,0 @@ -/- -Copyright (c) 2024 Christian Merten. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Christian Merten --/ -import Mathlib.RingTheory.Flat.Stability -import Mathlib.LinearAlgebra.TensorProduct.Tower - -/-! -# Flat algebras and flat ring homomorphisms - -We define flat algebras and flat ring homomorphisms. - -## Main definitions - -* `Algebra.Flat`: an `R`-algebra `S` is flat if it is flat as `R`-module -* `RingHom.Flat`: a ring homomorphism `f : R → S` is flat, if it makes `S` into a flat `R`-algebra - --/ - -universe u v w t - -open Function (Injective Surjective) - -open LinearMap (lsmul rTensor lTensor) - -open TensorProduct - -/-- An `R`-algebra `S` is flat if it is flat as `R`-module. -/ -class Algebra.Flat (R : Type u) (S : Type v) [CommRing R] [CommRing S] [Algebra R S] : Prop where - out : Module.Flat R S - -namespace Algebra.Flat - -attribute [instance] out - -instance self (R : Type u) [CommRing R] : Algebra.Flat R R where - out := Module.Flat.self R - -variable (R : Type u) (S : Type v) [CommRing R] [CommRing S] - -/-- If `T` is a flat `S`-algebra and `S` is a flat `R`-algebra, -then `T` is a flat `R`-algebra. -/ -theorem trans (T : Type w) [CommRing T] [Algebra R S] [Algebra R T] [Algebra S T] - [IsScalarTower R S T] [Algebra.Flat R S] [Algebra.Flat S T] : Algebra.Flat R T where - out := Module.Flat.trans R S T - -@[deprecated (since := "2024-11-08")] alias comp := trans - -/-- If `S` is a flat `R`-algebra and `T` is any `R`-algebra, -then `T ⊗[R] S` is a flat `T`-algebra. -/ -instance baseChange (T : Type w) [CommRing T] [Algebra R S] [Algebra R T] [Algebra.Flat R S] : - Algebra.Flat T (T ⊗[R] S) where - out := Module.Flat.baseChange R T S - -/-- A base change of a flat algebra is flat. -/ -theorem isBaseChange [Algebra R S] (R' : Type w) (S' : Type t) [CommRing R'] [CommRing S'] - [Algebra R R'] [Algebra S S'] [Algebra R' S'] [Algebra R S'] [IsScalarTower R R' S'] - [IsScalarTower R S S'] [h : IsPushout R S R' S'] [Algebra.Flat R R'] : Algebra.Flat S S' where - out := Module.Flat.isBaseChange R S R' S' h.out - -end Algebra.Flat - -/-- A ring homomorphism `f : R →+* S` is flat if `S` is flat as an `R` algebra. -/ -@[algebraize RingHom.Flat.out] -class RingHom.Flat {R : Type u} {S : Type v} [CommRing R] [CommRing S] (f : R →+* S) : Prop where - out : f.toAlgebra.Flat := by infer_instance - -namespace RingHom.Flat - -attribute [instance] out - -/-- The identity of a ring is flat. -/ -instance identity (R : Type u) [CommRing R] : RingHom.Flat (1 : R →+* R) where - -variable {R : Type u} {S : Type v} {T : Type w} [CommRing R] [CommRing S] [CommRing T] - (f : R →+* S) (g : S →+* T) - -/-- Composition of flat ring homomorphisms is flat. -/ -instance comp [RingHom.Flat f] [RingHom.Flat g] : RingHom.Flat (g.comp f) where - out := by - algebraize_only [f, g, g.comp f] - exact Algebra.Flat.trans R S T - -end RingHom.Flat diff --git a/Mathlib/RingTheory/RingHom/Flat.lean b/Mathlib/RingTheory/RingHom/Flat.lean new file mode 100644 index 0000000000000..ee70f2de16d6b --- /dev/null +++ b/Mathlib/RingTheory/RingHom/Flat.lean @@ -0,0 +1,71 @@ +/- +Copyright (c) 2024 Christian Merten. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Christian Merten +-/ +import Mathlib.RingTheory.Flat.Localization +import Mathlib.RingTheory.LocalProperties.Basic + +/-! +# Flat ring homomorphisms + +In this file we define flat ring homomorphisms and show their meta properties. + +-/ + +universe u v + +open TensorProduct + +/-- A ring homomorphism `f : R →+* S` is flat if `S` is flat as an `R` module. -/ +@[algebraize Module.Flat] +def RingHom.Flat {R : Type u} {S : Type v} [CommRing R] [CommRing S] (f : R →+* S) : Prop := + letI : Algebra R S := f.toAlgebra + Module.Flat R S + +namespace RingHom.Flat + +variable {R S T : Type*} [CommRing R] [CommRing S] [CommRing T] + +variable (R) in +/-- The identity of a ring is flat. -/ +lemma id : RingHom.Flat (RingHom.id R) := + Module.Flat.self R + +/-- Composition of flat ring homomorphisms is flat. -/ +lemma comp {f : R →+* S} {g : S →+* T} (hf : f.Flat) (hg : g.Flat) : Flat (g.comp f) := by + algebraize [f, g, (g.comp f)] + exact Module.Flat.trans R S T + +/-- Bijective ring maps are flat. -/ +lemma of_bijective {f : R →+* S} (hf : Function.Bijective f) : Flat f := by + algebraize [f] + exact Module.Flat.of_linearEquiv R R S (LinearEquiv.ofBijective (Algebra.linearMap R S) hf).symm + +lemma containsIdentities : ContainsIdentities Flat := id + +lemma stableUnderComposition : StableUnderComposition Flat := by + introv R hf hg + exact hf.comp hg + +lemma respectsIso : RespectsIso Flat := by + apply stableUnderComposition.respectsIso + introv + exact of_bijective e.bijective + +lemma isStableUnderBaseChange : IsStableUnderBaseChange Flat := by + apply IsStableUnderBaseChange.mk _ respectsIso + introv h + replace h : Module.Flat R T := by + rw [RingHom.Flat] at h; convert h; ext; simp_rw [Algebra.smul_def]; rfl + suffices Module.Flat S (S ⊗[R] T) by + rw [RingHom.Flat]; convert this; congr; ext; simp_rw [Algebra.smul_def]; rfl + exact inferInstance + +lemma holdsForLocalizationAway : HoldsForLocalizationAway Flat := by + introv R h + suffices Module.Flat R S by + rw [RingHom.Flat]; convert this; ext; simp_rw [Algebra.smul_def]; rfl + exact IsLocalization.flat _ (Submonoid.powers r) + +end RingHom.Flat From 37699c06515f88e75d2bb155457b3f8e456a5ec5 Mon Sep 17 00:00:00 2001 From: "Tristan F.-R." Date: Mon, 9 Dec 2024 08:25:51 +0000 Subject: [PATCH 580/829] docs(Combinatorics/Enumerative/Catalan): update theorem name in docs (#19792) Updates an outdated name for a theorem in the top-level module documentation. Co-authored-by: Tristan F. --- Mathlib/Combinatorics/Enumerative/Catalan.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Combinatorics/Enumerative/Catalan.lean b/Mathlib/Combinatorics/Enumerative/Catalan.lean index 21a578fab89bf..d7dd958d1d693 100644 --- a/Mathlib/Combinatorics/Enumerative/Catalan.lean +++ b/Mathlib/Combinatorics/Enumerative/Catalan.lean @@ -30,7 +30,7 @@ triangulations of convex polygons. * `catalan_eq_centralBinom_div`: The explicit formula for the Catalan number using the central binomial coefficient, `catalan n = Nat.centralBinom n / (n + 1)`. -* `treesOfNodesEq_card_eq_catalan`: The number of binary trees with `n` internal nodes +* `treesOfNumNodesEq_card_eq_catalan`: The number of binary trees with `n` internal nodes is `catalan n` ## Implementation details From 15fde81e3d3ffe0e30031cbdf8a3fb3e23d2a927 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 9 Dec 2024 08:25:53 +0000 Subject: [PATCH 581/829] =?UTF-8?q?feat(Order/SuccPred/CompleteLinearOrder?= =?UTF-8?q?):=20`sSup=20(Iio=20x)=20=3D=20x=20=E2=86=94=20IsSuccPrelimit?= =?UTF-8?q?=20x`=20(#19800)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- .../Order/SuccPred/CompleteLinearOrder.lean | 28 ++++++++++++++++++- 1 file changed, 27 insertions(+), 1 deletion(-) diff --git a/Mathlib/Order/SuccPred/CompleteLinearOrder.lean b/Mathlib/Order/SuccPred/CompleteLinearOrder.lean index c8ab06d7b669a..eea7681180495 100644 --- a/Mathlib/Order/SuccPred/CompleteLinearOrder.lean +++ b/Mathlib/Order/SuccPred/CompleteLinearOrder.lean @@ -11,7 +11,7 @@ import Mathlib.Order.SuccPred.Limit -/ -open Order +open Order Set variable {ι α : Type*} @@ -144,6 +144,32 @@ lemma IsLUB.exists_of_not_isSuccPrelimit (hf : IsLUB (Set.range f) x) (hx : ¬ I @[deprecated IsLUB.exists_of_not_isSuccPrelimit (since := "2024-09-05")] alias IsLUB.exists_of_not_isSuccLimit := IsLUB.exists_of_not_isSuccPrelimit +theorem Order.IsSuccPrelimit.sSup_Iio (h : IsSuccPrelimit x) : sSup (Iio x) = x := by + obtain rfl | hx := eq_bot_or_bot_lt x + · simp + · refine (csSup_le ⟨⊥, hx⟩ fun a ha ↦ ha.le).antisymm <| le_of_forall_lt fun a ha ↦ ?_ + rw [lt_csSup_iff' bddAbove_Iio] + obtain ⟨b, hb', hb⟩ := (not_covBy_iff ha).1 (h a) + use b, hb + +theorem Order.IsSuccPrelimit.iSup_Iio (h : IsSuccPrelimit x) : ⨆ a : Iio x, a.1 = x := by + rw [← sSup_eq_iSup', h.sSup_Iio] + +theorem Order.IsSuccLimit.sSup_Iio (h : IsSuccLimit x) : sSup (Iio x) = x := + h.isSuccPrelimit.sSup_Iio + +theorem Order.IsSuccLimit.iSup_Iio (h : IsSuccLimit x) : ⨆ a : Iio x, a.1 = x := + h.isSuccPrelimit.iSup_Iio + +theorem sSup_Iio_eq_self_iff_isSuccPrelimit : sSup (Iio x) = x ↔ IsSuccPrelimit x := by + refine ⟨fun h ↦ ?_, IsSuccPrelimit.sSup_Iio⟩ + by_contra hx + rw [← h] at hx + simpa [h] using csSup_mem_of_not_isSuccPrelimit' bddAbove_Iio hx + +theorem iSup_Iio_eq_self_iff_isSuccPrelimit : ⨆ a : Iio x, a.1 = x ↔ IsSuccPrelimit x := by + rw [← sSup_eq_iSup', sSup_Iio_eq_self_iff_isSuccPrelimit] + end ConditionallyCompleteLinearOrderBot section CompleteLinearOrder From b670b3cab336b8f2080947d16d2b2da982357f9c Mon Sep 17 00:00:00 2001 From: sgouezel Date: Mon, 9 Dec 2024 08:25:55 +0000 Subject: [PATCH 582/829] chore: cleanup of manifold files (#19807) As a preparation for #19696 --- Mathlib/Analysis/Calculus/ContDiff/Defs.lean | 7 +------ .../Calculus/ContDiff/FTaylorSeries.lean | 18 +++++++++-------- .../Geometry/Manifold/Algebra/LieGroup.lean | 12 +++++------ Mathlib/Geometry/Manifold/ChartedSpace.lean | 4 ++++ Mathlib/Geometry/Manifold/ContMDiff/Defs.lean | 20 +++++++++++++++++++ Mathlib/Geometry/Manifold/Diffeomorph.lean | 4 +--- Mathlib/Geometry/Manifold/Instances/Real.lean | 8 +++++--- .../Manifold/MFDeriv/NormedSpace.lean | 10 +++++----- .../Manifold/SmoothManifoldWithCorners.lean | 8 +++++--- 9 files changed, 57 insertions(+), 34 deletions(-) diff --git a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean index 58258d4bbd896..6f8cc0631548c 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean @@ -107,12 +107,7 @@ attribute [local instance 1001] open Set Fin Filter Function -/-- Smoothness exponent for analytic functions. -/ -scoped [ContDiff] notation3 "ω" => (⊤ : WithTop ℕ∞) -/-- Smoothness exponent for infinitely differentiable functions. -/ -scoped [ContDiff] notation3 "∞" => ((⊤ : ℕ∞) : WithTop ℕ∞) - -open ContDiff +open scoped ContDiff universe u uE uF uG uX diff --git a/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean b/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean index 95ea975bb88e6..a66bfc2242b36 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean @@ -98,10 +98,7 @@ In this file, we denote `⊤ : ℕ∞` with `∞`. noncomputable section -open scoped Classical -open ENat NNReal Topology Filter - -local notation "∞" => (⊤ : ℕ∞) +open ENat NNReal Topology Filter Set Fin Filter Function /- Porting note: These lines are not required in Mathlib4. @@ -109,7 +106,12 @@ attribute [local instance 1001] NormedAddCommGroup.toAddCommGroup NormedSpace.toModule' AddCommGroup.toAddCommMonoid -/ -open Set Fin Filter Function +/-- Smoothness exponent for analytic functions. -/ +scoped [ContDiff] notation3 "ω" => (⊤ : WithTop ℕ∞) +/-- Smoothness exponent for infinitely differentiable functions. -/ +scoped [ContDiff] notation3 "∞" => ((⊤ : ℕ∞) : WithTop ℕ∞) + +open scoped ContDiff universe u uE uF @@ -129,7 +131,7 @@ Notice that `p` does not sum up to `f` on the diagonal (`FormalMultilinearSeries structure HasFTaylorSeriesUpToOn (n : WithTop ℕ∞) (f : E → F) (p : E → FormalMultilinearSeries 𝕜 E F) (s : Set E) : Prop where zero_eq : ∀ x ∈ s, (p x 0).curry0 = f x - protected fderivWithin : ∀ m : ℕ, (m : ℕ∞) < n → ∀ x ∈ s, + protected fderivWithin : ∀ m : ℕ, m < n → ∀ x ∈ s, HasFDerivWithinAt (p · m) (p x m.succ).curryLeft s x cont : ∀ m : ℕ, m ≤ n → ContinuousOn (p · m) s @@ -609,7 +611,7 @@ theorem HasFTaylorSeriesUpToOn.eq_iteratedFDerivWithin_of_uniqueDiffOn (hx : x ∈ s) : p x m = iteratedFDerivWithin 𝕜 m f s x := by induction' m with m IH generalizing x · rw [h.zero_eq' hx, iteratedFDerivWithin_zero_eq_comp]; rfl - · have A : (m : ℕ∞) < n := lt_of_lt_of_le (mod_cast lt_add_one m) hmn + · have A : m < n := lt_of_lt_of_le (mod_cast lt_add_one m) hmn have : HasFDerivWithinAt (fun y : E => iteratedFDerivWithin 𝕜 m f s y) (ContinuousMultilinearMap.curryLeft (p x (Nat.succ m))) s x := @@ -807,7 +809,7 @@ theorem iteratedFDerivWithin_univ {n : ℕ} : rw [iteratedFDeriv_succ_apply_left, iteratedFDerivWithin_succ_apply_left, IH, fderivWithin_univ] theorem HasFTaylorSeriesUpTo.eq_iteratedFDeriv - (h : HasFTaylorSeriesUpTo n f p) {m : ℕ} (hmn : (m : ℕ∞) ≤ n) (x : E) : + (h : HasFTaylorSeriesUpTo n f p) {m : ℕ} (hmn : m ≤ n) (x : E) : p x m = iteratedFDeriv 𝕜 m f x := by rw [← iteratedFDerivWithin_univ] rw [← hasFTaylorSeriesUpToOn_univ_iff] at h diff --git a/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean b/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean index cae9a80c72889..2997b5c92169b 100644 --- a/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean +++ b/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean @@ -211,7 +211,7 @@ instance {𝕜 : Type*} [NontriviallyNormedField 𝕜] : SmoothInv₀ 𝓘(𝕜) exact contDiffAt_inv 𝕜 hx variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [TopologicalSpace H] {E : Type*} - [NormedAddCommGroup E] [NormedSpace 𝕜 E] (I : ModelWithCorners 𝕜 E H) {G : Type*} + [NormedAddCommGroup E] [NormedSpace 𝕜 E] {I : ModelWithCorners 𝕜 E H} {G : Type*} [TopologicalSpace G] [ChartedSpace H G] [Inv G] [Zero G] [SmoothInv₀ I G] {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} {M : Type*} [TopologicalSpace M] [ChartedSpace H' M] @@ -227,23 +227,23 @@ include I in This is not an instance for technical reasons, see note [Design choices about smooth algebraic structures]. -/ theorem hasContinuousInv₀_of_hasSmoothInv₀ : HasContinuousInv₀ G := - { continuousAt_inv₀ := fun _ hx ↦ (contMDiffAt_inv₀ I hx).continuousAt } + { continuousAt_inv₀ := fun _ hx ↦ (contMDiffAt_inv₀ (I := I) hx).continuousAt } theorem contMDiffOn_inv₀ : ContMDiffOn I I ⊤ (Inv.inv : G → G) {0}ᶜ := fun _x hx => - (contMDiffAt_inv₀ I hx).contMDiffWithinAt + (contMDiffAt_inv₀ hx).contMDiffWithinAt @[deprecated (since := "2024-11-21")] alias smoothOn_inv₀ := contMDiffOn_inv₀ @[deprecated (since := "2024-11-21")] alias SmoothOn_inv₀ := contMDiffOn_inv₀ -variable {I} {s : Set M} {a : M} +variable {s : Set M} {a : M} theorem ContMDiffWithinAt.inv₀ (hf : ContMDiffWithinAt I' I n f s a) (ha : f a ≠ 0) : ContMDiffWithinAt I' I n (fun x => (f x)⁻¹) s a := - ((contMDiffAt_inv₀ I ha).of_le le_top).comp_contMDiffWithinAt a hf + ((contMDiffAt_inv₀ ha).of_le le_top).comp_contMDiffWithinAt a hf theorem ContMDiffAt.inv₀ (hf : ContMDiffAt I' I n f a) (ha : f a ≠ 0) : ContMDiffAt I' I n (fun x ↦ (f x)⁻¹) a := - ((contMDiffAt_inv₀ I ha).of_le le_top).comp a hf + ((contMDiffAt_inv₀ ha).of_le le_top).comp a hf theorem ContMDiff.inv₀ (hf : ContMDiff I' I n f) (h0 : ∀ x, f x ≠ 0) : ContMDiff I' I n (fun x ↦ (f x)⁻¹) := diff --git a/Mathlib/Geometry/Manifold/ChartedSpace.lean b/Mathlib/Geometry/Manifold/ChartedSpace.lean index 93c0039fd5c2f..2c3d9ef063c40 100644 --- a/Mathlib/Geometry/Manifold/ChartedSpace.lean +++ b/Mathlib/Geometry/Manifold/ChartedSpace.lean @@ -1051,6 +1051,10 @@ theorem StructureGroupoid.mem_maximalAtlas_of_mem_groupoid {f : PartialHomeomorp rintro e (rfl : e = PartialHomeomorph.refl H) exact ⟨G.trans (G.symm hf) G.id_mem, G.trans (G.symm G.id_mem) hf⟩ +theorem StructureGroupoid.maximalAtlas_mono {G G' : StructureGroupoid H} (h : G ≤ G') : + G.maximalAtlas M ⊆ G'.maximalAtlas M := + fun _ he e' he' ↦ ⟨h (he e' he').1, h (he e' he').2⟩ + end MaximalAtlas section Singleton diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean index 112fab89cc177..3ab6587a89afc 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean @@ -479,6 +479,21 @@ theorem contMDiffOn_iff : convert hdiff x (f x) (extChartAt I x x) (by simp only [hx, mfld_simps]) using 1 mfld_set_tac +/-- zero-smoothness on a set is equivalent to continuity on this set. -/ +theorem contMDiffOn_zero_iff : + ContMDiffOn I I' 0 f s ↔ ContinuousOn f s := by + rw [contMDiffOn_iff] + refine ⟨fun h ↦ h.1, fun h ↦ ⟨h, ?_⟩⟩ + intro x y + rw [show ((0 : ℕ∞) : WithTop ℕ∞) = 0 by rfl, contDiffOn_zero] + apply (continuousOn_extChartAt _).comp + · apply h.comp ((continuousOn_extChartAt_symm _).mono inter_subset_left) (fun z hz ↦ ?_) + simp only [preimage_inter, mem_inter_iff, mem_preimage] at hz + exact hz.2.1 + · intro z hz + simp only [preimage_inter, mem_inter_iff, mem_preimage] at hz + exact hz.2.2 + /-- One can reformulate smoothness on a set as continuity on this set, and smoothness in any extended chart in the target. -/ theorem contMDiffOn_iff_target : @@ -525,6 +540,11 @@ theorem contMDiff_iff_target : @[deprecated (since := "2024-11-20")] alias smooth_iff_target := contMDiff_iff_target +/-- zero-smoothness is equivalent to continuity. -/ +theorem contMDiff_zero_iff : + ContMDiff I I' 0 f ↔ Continuous f := by + rw [← contMDiffOn_univ, continuous_iff_continuousOn_univ, contMDiffOn_zero_iff] + end SmoothManifoldWithCorners /-! ### Deducing smoothness from smoothness one step beyond -/ diff --git a/Mathlib/Geometry/Manifold/Diffeomorph.lean b/Mathlib/Geometry/Manifold/Diffeomorph.lean index 6a6dd316cf875..ced591a0a2f5f 100644 --- a/Mathlib/Geometry/Manifold/Diffeomorph.lean +++ b/Mathlib/Geometry/Manifold/Diffeomorph.lean @@ -79,9 +79,7 @@ scoped[Manifold] notation M " ≃ₘ^" n:1000 "⟮" I ", " J "⟯ " N => Diffeom scoped[Manifold] notation M " ≃ₘ⟮" I ", " J "⟯ " N => Diffeomorph I J M N ⊤ /-- `n`-times continuously differentiable diffeomorphism between `E` and `E'`. -/ -scoped[Manifold] - notation E " ≃ₘ^" n:1000 "[" 𝕜 "] " E' => - Diffeomorph (modelWithCornersSelf 𝕜 E) (modelWithCornersSelf 𝕜 E') E E' n +scoped[Manifold] notation E " ≃ₘ^" n:1000 "[" 𝕜 "] " E' => Diffeomorph 𝓘(𝕜, E) 𝓘(𝕜, E') E E' n /-- Infinitely differentiable diffeomorphism between `E` and `E'`. -/ scoped[Manifold] diff --git a/Mathlib/Geometry/Manifold/Instances/Real.lean b/Mathlib/Geometry/Manifold/Instances/Real.lean index 8e667f22b2b24..7f10c097d74fe 100644 --- a/Mathlib/Geometry/Manifold/Instances/Real.lean +++ b/Mathlib/Geometry/Manifold/Instances/Real.lean @@ -14,16 +14,18 @@ or with boundary or with corners. As a concrete example, we construct explicitly boundary structure on the real interval `[x, y]`. More specifically, we introduce -* `ModelWithCorners ℝ (EuclideanSpace ℝ (Fin n)) (EuclideanHalfSpace n)` for the model space +* `modelWithCornersEuclideanHalfSpace n : + ModelWithCorners ℝ (EuclideanSpace ℝ (Fin n)) (EuclideanHalfSpace n)` for the model space used to define `n`-dimensional real manifolds with boundary -* `ModelWithCorners ℝ (EuclideanSpace ℝ (Fin n)) (EuclideanQuadrant n)` for the model space used +* `modelWithCornersEuclideanQuadrant n : + ModelWithCorners ℝ (EuclideanSpace ℝ (Fin n)) (EuclideanQuadrant n)` for the model space used to define `n`-dimensional real manifolds with corners ## Notations In the locale `Manifold`, we introduce the notations * `𝓡 n` for the identity model with corners on `EuclideanSpace ℝ (Fin n)` -* `𝓡∂ n` for `ModelWithCorners ℝ (EuclideanSpace ℝ (Fin n)) (EuclideanHalfSpace n)`. +* `𝓡∂ n` for `modelWithCornersEuclideanHalfSpace n`. For instance, if a manifold `M` is boundaryless, smooth and modelled on `EuclideanSpace ℝ (Fin m)`, and `N` is smooth with boundary modelled on `EuclideanHalfSpace n`, and `f : M → N` is a smooth diff --git a/Mathlib/Geometry/Manifold/MFDeriv/NormedSpace.lean b/Mathlib/Geometry/Manifold/MFDeriv/NormedSpace.lean index b30e07d620aa0..2e727c4ef5022 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/NormedSpace.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/NormedSpace.lean @@ -79,7 +79,7 @@ theorem Differentiable.comp_mdifferentiable {g : F → F'} {f : M → F} (hg : D end Module -/-! ### Linear maps between normed spaces are smooth -/ +/-! ### Linear maps between normed spaces are differentiable -/ #adaptation_note /-- @@ -189,8 +189,8 @@ theorem MDifferentiable.clm_comp {g : M → F₁ →L[𝕜] F₃} {f : M → F (hg : MDifferentiable I 𝓘(𝕜, F₁ →L[𝕜] F₃) g) (hf : MDifferentiable I 𝓘(𝕜, F₂ →L[𝕜] F₁) f) : MDifferentiable I 𝓘(𝕜, F₂ →L[𝕜] F₃) fun x => (g x).comp (f x) := fun x => (hg x).clm_comp (hf x) -/-- Applying a linear map to a vector is smooth within a set. Version in vector spaces. For a -version in nontrivial vector bundles, see `MDifferentiableWithinAt.clm_apply_of_inCoordinates`. -/ +/-- Applying a linear map to a vector is differentiable within a set. Version in vector spaces. For +a version in nontrivial vector bundles, see `MDifferentiableWithinAt.clm_apply_of_inCoordinates`. -/ theorem MDifferentiableWithinAt.clm_apply {g : M → F₁ →L[𝕜] F₂} {f : M → F₁} {s : Set M} {x : M} (hg : MDifferentiableWithinAt I 𝓘(𝕜, F₁ →L[𝕜] F₂) g s x) (hf : MDifferentiableWithinAt I 𝓘(𝕜, F₁) f s x) : @@ -201,7 +201,7 @@ theorem MDifferentiableWithinAt.clm_apply {g : M → F₁ →L[𝕜] F₂} {f : exact differentiable_fst.clm_apply differentiable_snd) (hg.prod_mk_space hf) (by simp_rw [mapsTo_univ]) -/-- Applying a linear map to a vector is smooth. Version in vector spaces. For a +/-- Applying a linear map to a vector is differentiable. Version in vector spaces. For a version in nontrivial vector bundles, see `MDifferentiableAt.clm_apply_of_inCoordinates`. -/ theorem MDifferentiableAt.clm_apply {g : M → F₁ →L[𝕜] F₂} {f : M → F₁} {x : M} (hg : MDifferentiableAt I 𝓘(𝕜, F₁ →L[𝕜] F₂) g x) (hf : MDifferentiableAt I 𝓘(𝕜, F₁) f x) : @@ -282,7 +282,7 @@ theorem MDifferentiable.clm_prodMap {g : M → F₁ →L[𝕜] F₃} {f : M → MDifferentiable I 𝓘(𝕜, F₁ × F₂ →L[𝕜] F₃ × F₄) fun x => (g x).prodMap (f x) := fun x => (hg x).clm_prodMap (hf x) -/-! ### Smoothness of scalar multiplication -/ +/-! ### Differentiability of scalar multiplication -/ variable {V : Type*} [NormedAddCommGroup V] [NormedSpace 𝕜 V] diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index 7511ccb7ad482..9fcbaff7d4fac 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -48,12 +48,14 @@ but add these assumptions later as needed. (Quite a few results still do not req `extChartAt I x` is the canonical such partial equiv around `x`. As specific examples of models with corners, we define (in `Geometry.Manifold.Instances.Real`) -* `modelWithCornersSelf ℝ (EuclideanSpace ℝ (Fin n))` for the model space used to define +* `modelWithCornersEuclideanHalfSpace n : + ModelWithCorners ℝ (EuclideanSpace ℝ (Fin n)) (EuclideanHalfSpace n)` for the model space `n`-dimensional real manifolds without boundary (with notation `𝓡 n` in the locale `Manifold`) * `ModelWithCorners ℝ (EuclideanSpace ℝ (Fin n)) (EuclideanHalfSpace n)` for the model space used to define `n`-dimensional real manifolds with boundary (with notation `𝓡∂ n` in the locale `Manifold`) -* `ModelWithCorners ℝ (EuclideanSpace ℝ (Fin n)) (EuclideanQuadrant n)` for the model space used +* `modelWithCornersEuclideanQuadrant n : + ModelWithCorners ℝ (EuclideanSpace ℝ (Fin n)) (EuclideanQuadrant n)` for the model space used to define `n`-dimensional real manifolds with corners With these definitions at hand, to invoke an `n`-dimensional real manifold without boundary, @@ -369,7 +371,7 @@ protected theorem secondCountableTopology [SecondCountableTopology E] (I : Model I.isClosedEmbedding.isEmbedding.secondCountableTopology include I in -/-- Every smooth manifold is a Fréchet space (T1 space) -- regardless of whether it is +/-- Every manifold is a Fréchet space (T1 space) -- regardless of whether it is Hausdorff. -/ protected theorem t1Space (M : Type*) [TopologicalSpace M] [ChartedSpace H M] : T1Space M := by have : T2Space H := I.isClosedEmbedding.toIsEmbedding.t2Space From ff66079481b64d88e93f95608d309579e12cdd0c Mon Sep 17 00:00:00 2001 From: Pieter Cuijpers Date: Mon, 9 Dec 2024 08:25:56 +0000 Subject: [PATCH 583/829] chore: generalize OrderMonoidIso to OrderMulIso (#19810) An OrderMonoidIso actually does not depend on the fact that the type has a One, so the assumption [MulOneClass] can be relaxed to [Mul] in the definition, and similar for the AddOrderMonoidIso. For the time being, let's not rename the definition yet, just relax it. (I need this for my work on Quantales, as a OrderMulIso is also a QuantaleIso, but Quantales do not necessarily have a One). --- Mathlib/Algebra/Order/Hom/Monoid.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Algebra/Order/Hom/Monoid.lean b/Mathlib/Algebra/Order/Hom/Monoid.lean index b40dacaccb13e..d7ee5915f5f32 100644 --- a/Mathlib/Algebra/Order/Hom/Monoid.lean +++ b/Mathlib/Algebra/Order/Hom/Monoid.lean @@ -90,8 +90,8 @@ structure. When possible, instead of parametrizing results over `(f : α ≃+o β)`, you should parametrize over `(F : Type*) [FunLike F M N] [AddEquivClass F M N] [OrderIsoClass F M N] (f : F)`. -/ -structure OrderAddMonoidIso (α β : Type*) [Preorder α] [Preorder β] [AddZeroClass α] - [AddZeroClass β] extends α ≃+ β where +structure OrderAddMonoidIso (α β : Type*) [Preorder α] [Preorder β] [Add α] [Add β] + extends α ≃+ β where /-- An `OrderAddMonoidIso` respects `≤`. -/ map_le_map_iff' {a b : α} : toFun a ≤ toFun b ↔ a ≤ b @@ -146,8 +146,8 @@ When possible, instead of parametrizing results over `(f : α ≃*o β)`, you should parametrize over `(F : Type*) [FunLike F M N] [MulEquivClass F M N] [OrderIsoClass F M N] (f : F)`. -/ @[to_additive] -structure OrderMonoidIso (α β : Type*) [Preorder α] [Preorder β] [MulOneClass α] - [MulOneClass β] extends α ≃* β where +structure OrderMonoidIso (α β : Type*) [Preorder α] [Preorder β] [Mul α] [Mul β] + extends α ≃* β where /-- An `OrderMonoidIso` respects `≤`. -/ map_le_map_iff' {a b : α} : toFun a ≤ toFun b ↔ a ≤ b @@ -513,8 +513,8 @@ namespace OrderMonoidIso section Preorder -variable [Preorder α] [Preorder β] [Preorder γ] [Preorder δ] [MulOneClass α] [MulOneClass β] - [MulOneClass γ] [MulOneClass δ] {f g : α ≃*o β} +variable [Preorder α] [Preorder β] [Preorder γ] [Preorder δ] [Mul α] [Mul β] + [Mul γ] [Mul δ] {f g : α ≃*o β} @[to_additive] instance : EquivLike (α ≃*o β) α β where From d9df436d7f9d6e2eb0aad05d429d2fc53d4e301d Mon Sep 17 00:00:00 2001 From: Pieter Cuijpers Date: Mon, 9 Dec 2024 08:25:58 +0000 Subject: [PATCH 584/829] =?UTF-8?q?feat(Quantale):=20`=E2=8A=A5=20*=20x=20?= =?UTF-8?q?=3D=20=E2=8A=A5`=20(#19811)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The bottom element of a quantale is absorbing for multiplication As this is a basic truth for quantales that one practically always will use, I think it belongs in the main Quantale file. --- Mathlib/Algebra/Order/Quantale.lean | 17 +++++++++++++++++ 1 file changed, 17 insertions(+) diff --git a/Mathlib/Algebra/Order/Quantale.lean b/Mathlib/Algebra/Order/Quantale.lean index 96bfa1624aa35..13108eac3152d 100644 --- a/Mathlib/Algebra/Order/Quantale.lean +++ b/Mathlib/Algebra/Order/Quantale.lean @@ -176,4 +176,21 @@ theorem rightMulResiduation_le_iff_mul_le : x ≤ y ⇨ᵣ z ↔ y * x ≤ z whe iSup_le_iff, implies_true] mpr h1 := le_sSup h1 +section Zero + +variable {α : Type*} [Semigroup α] [CompleteLattice α] [IsQuantale α] +variable {x : α} + +@[to_additive (attr := simp)] +theorem bot_mul : ⊥ * x = ⊥ := by + rw [← sSup_empty, sSup_mul_distrib] + simp only [Set.mem_empty_iff_false, not_false_eq_true, iSup_neg, iSup_bot, sSup_empty] + +@[to_additive (attr := simp)] +theorem mul_bot : x * ⊥ = ⊥ := by + rw [← sSup_empty, mul_sSup_distrib] + simp only [Set.mem_empty_iff_false, not_false_eq_true, iSup_neg, iSup_bot, sSup_empty] + +end Zero + end IsQuantale From f064040f50bee5be6d2bb130c7fe7917db0df441 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Mon, 9 Dec 2024 09:03:54 +0000 Subject: [PATCH 585/829] chore(Algebra/Ring): generalize Injective.isDomain (#19768) --- .../CanonicallyOrderedCommSemiringTwoMul.lean | 3 +- Mathlib/Algebra/GroupWithZero/InjSurj.lean | 37 ++++++++------- Mathlib/Algebra/GroupWithZero/NeZero.lean | 4 +- Mathlib/Algebra/Order/Ring/InjSurj.lean | 2 +- Mathlib/Algebra/Ring/Basic.lean | 46 +++++++------------ Mathlib/Algebra/Ring/Hom/Basic.lean | 10 ++-- Mathlib/Data/Complex/Basic.lean | 2 +- Mathlib/Logic/Equiv/TransferInstance.lean | 6 +-- Mathlib/RingTheory/Ideal/Prime.lean | 1 - .../RingTheory/LocalRing/RingHom/Basic.lean | 2 +- .../Polynomial/IntegralNormalization.lean | 1 - 11 files changed, 52 insertions(+), 62 deletions(-) diff --git a/Counterexamples/CanonicallyOrderedCommSemiringTwoMul.lean b/Counterexamples/CanonicallyOrderedCommSemiringTwoMul.lean index 832f9e8bef97d..b2961dcf17622 100644 --- a/Counterexamples/CanonicallyOrderedCommSemiringTwoMul.lean +++ b/Counterexamples/CanonicallyOrderedCommSemiringTwoMul.lean @@ -101,8 +101,7 @@ theorem mul_lt_mul_of_pos_right : ∀ a b c : ℕ × ZMod 2, a < b → 0 < c → instance socsN2 : StrictOrderedCommSemiring (ℕ × ZMod 2) := { Nxzmod2.csrN21, (inferInstance : PartialOrder (ℕ × ZMod 2)), (inferInstance : CommSemiring (ℕ × ZMod 2)), - pullback_nonzero Prod.fst Prod.fst_zero - Prod.fst_one with + domain_nontrivial Prod.fst Prod.fst_zero Prod.fst_one with add_le_add_left := add_le_add_left le_of_add_le_add_left := le_of_add_le_add_left zero_le_one := zero_le_one diff --git a/Mathlib/Algebra/GroupWithZero/InjSurj.lean b/Mathlib/Algebra/GroupWithZero/InjSurj.lean index 75dc801b0bbb6..fb65d3fb24e9c 100644 --- a/Mathlib/Algebra/GroupWithZero/InjSurj.lean +++ b/Mathlib/Algebra/GroupWithZero/InjSurj.lean @@ -51,25 +51,30 @@ variable [Mul M₀] [Zero M₀] [Mul M₀'] [Zero M₀'] include hf zero mul /-- Pull back a `NoZeroDivisors` instance along an injective function. -/ -protected theorem Function.Injective.noZeroDivisors [NoZeroDivisors M₀'] : NoZeroDivisors M₀ := - { eq_zero_or_eq_zero_of_mul_eq_zero := fun {a b} H ↦ - have : f a * f b = 0 := by rw [← mul, H, zero] - (eq_zero_or_eq_zero_of_mul_eq_zero this).imp - (fun H => hf <| by rwa [zero]) fun H => hf <| by rwa [zero] } +protected theorem Function.Injective.noZeroDivisors [NoZeroDivisors M₀'] : NoZeroDivisors M₀ where + eq_zero_or_eq_zero_of_mul_eq_zero {a b} H := + have : f a * f b = 0 := by rw [← mul, H, zero] + (eq_zero_or_eq_zero_of_mul_eq_zero this).imp + (fun H ↦ hf <| by rwa [zero]) fun H ↦ hf <| by rwa [zero] protected theorem Function.Injective.isLeftCancelMulZero - [IsLeftCancelMulZero M₀'] : IsLeftCancelMulZero M₀ := - { mul_left_cancel_of_ne_zero := fun Hne He => by - have := congr_arg f He - rw [mul, mul] at this - exact hf (mul_left_cancel₀ (fun Hfa => Hne <| hf <| by rw [Hfa, zero]) this) } + [IsLeftCancelMulZero M₀'] : IsLeftCancelMulZero M₀ where + mul_left_cancel_of_ne_zero Hne He := by + have := congr_arg f He + rw [mul, mul] at this + exact hf (mul_left_cancel₀ (fun Hfa => Hne <| hf <| by rw [Hfa, zero]) this) protected theorem Function.Injective.isRightCancelMulZero - [IsRightCancelMulZero M₀'] : IsRightCancelMulZero M₀ := - { mul_right_cancel_of_ne_zero := fun Hne He => by - have := congr_arg f He - rw [mul, mul] at this - exact hf (mul_right_cancel₀ (fun Hfa => Hne <| hf <| by rw [Hfa, zero]) this) } + [IsRightCancelMulZero M₀'] : IsRightCancelMulZero M₀ where + mul_right_cancel_of_ne_zero Hne He := by + have := congr_arg f He + rw [mul, mul] at this + exact hf (mul_right_cancel₀ (fun Hfa => Hne <| hf <| by rw [Hfa, zero]) this) + +protected theorem Function.Injective.isCancelMulZero + [IsCancelMulZero M₀'] : IsCancelMulZero M₀ where + __ := hf.isLeftCancelMulZero f zero mul + __ := hf.isRightCancelMulZero f zero mul end NoZeroDivisors @@ -192,7 +197,7 @@ protected abbrev Function.Injective.groupWithZero [Zero G₀'] [Mul G₀'] [One (zpow : ∀ (x) (n : ℤ), f (x ^ n) = f x ^ n) : GroupWithZero G₀' := { hf.monoidWithZero f zero one mul npow, hf.divInvMonoid f one mul inv div npow zpow, - pullback_nonzero f zero one with + domain_nontrivial f zero one with inv_zero := hf <| by rw [inv, zero, inv_zero], mul_inv_cancel := fun x hx => hf <| by rw [one, mul, inv, mul_inv_cancel₀ ((hf.ne_iff' zero).2 hx)] } diff --git a/Mathlib/Algebra/GroupWithZero/NeZero.lean b/Mathlib/Algebra/GroupWithZero/NeZero.lean index 5f9e303a359d3..0d8a9a7bdb064 100644 --- a/Mathlib/Algebra/GroupWithZero/NeZero.lean +++ b/Mathlib/Algebra/GroupWithZero/NeZero.lean @@ -31,12 +31,14 @@ instance NeZero.one : NeZero (1 : M₀) := ⟨by _ = y := by rw [one_mul]⟩ /-- Pullback a `Nontrivial` instance along a function sending `0` to `0` and `1` to `1`. -/ -theorem pullback_nonzero [Zero M₀'] [One M₀'] (f : M₀' → M₀) (zero : f 0 = 0) (one : f 1 = 1) : +theorem domain_nontrivial [Zero M₀'] [One M₀'] (f : M₀' → M₀) (zero : f 0 = 0) (one : f 1 = 1) : Nontrivial M₀' := ⟨⟨0, 1, mt (congr_arg f) <| by rw [zero, one] exact zero_ne_one⟩⟩ +@[deprecated (since := "2024-12-07")] alias pullback_nonzero := domain_nontrivial + section GroupWithZero variable {G₀ : Type*} [GroupWithZero G₀] {a : G₀} diff --git a/Mathlib/Algebra/Order/Ring/InjSurj.lean b/Mathlib/Algebra/Order/Ring/InjSurj.lean index f8f659b32ebbc..1942166c409a7 100644 --- a/Mathlib/Algebra/Order/Ring/InjSurj.lean +++ b/Mathlib/Algebra/Order/Ring/InjSurj.lean @@ -74,7 +74,7 @@ protected abbrev strictOrderedSemiring [StrictOrderedSemiring α] (zero : f 0 = (natCast : ∀ n : ℕ, f n = n) : StrictOrderedSemiring β where toSemiring := hf.semiring f zero one add mul nsmul npow natCast __ := hf.orderedCancelAddCommMonoid f zero add (swap nsmul) - __ := pullback_nonzero f zero one + __ := domain_nontrivial f zero one __ := hf.orderedSemiring f zero one add mul nsmul npow natCast mul_lt_mul_of_pos_left a b c h hc := show f (c * a) < f (c * b) by simpa only [mul, zero] using mul_lt_mul_of_pos_left ‹f a < f b› (by rwa [← zero]) diff --git a/Mathlib/Algebra/Ring/Basic.lean b/Mathlib/Algebra/Ring/Basic.lean index fd10db100aadd..1682f8c413b9e 100644 --- a/Mathlib/Algebra/Ring/Basic.lean +++ b/Mathlib/Algebra/Ring/Basic.lean @@ -124,37 +124,25 @@ section NoZeroDivisors variable (α) -lemma IsLeftCancelMulZero.to_noZeroDivisors [NonUnitalNonAssocRing α] [IsLeftCancelMulZero α] : - NoZeroDivisors α := - { eq_zero_or_eq_zero_of_mul_eq_zero := fun {x y} h ↦ by - by_cases hx : x = 0 - { left - exact hx } - { right - rw [← sub_zero (x * y), ← mul_zero x, ← mul_sub] at h - have := (IsLeftCancelMulZero.mul_left_cancel_of_ne_zero) hx h - rwa [sub_zero] at this } } - -lemma IsRightCancelMulZero.to_noZeroDivisors [NonUnitalNonAssocRing α] [IsRightCancelMulZero α] : - NoZeroDivisors α := - { eq_zero_or_eq_zero_of_mul_eq_zero := fun {x y} h ↦ by - by_cases hy : y = 0 - { right - exact hy } - { left - rw [← sub_zero (x * y), ← zero_mul y, ← sub_mul] at h - have := (IsRightCancelMulZero.mul_right_cancel_of_ne_zero) hy h - rwa [sub_zero] at this } } +lemma IsLeftCancelMulZero.to_noZeroDivisors [NonUnitalNonAssocSemiring α] + [IsLeftCancelMulZero α] : NoZeroDivisors α where + eq_zero_or_eq_zero_of_mul_eq_zero {x _} h := + or_iff_not_imp_left.mpr fun ne ↦ mul_left_cancel₀ ne ((mul_zero x).symm ▸ h) + +lemma IsRightCancelMulZero.to_noZeroDivisors [NonUnitalNonAssocSemiring α] + [IsRightCancelMulZero α] : NoZeroDivisors α where + eq_zero_or_eq_zero_of_mul_eq_zero {_ y} h := + or_iff_not_imp_right.mpr fun ne ↦ mul_right_cancel₀ ne ((zero_mul y).symm ▸ h) instance (priority := 100) NoZeroDivisors.to_isCancelMulZero [NonUnitalNonAssocRing α] [NoZeroDivisors α] : - IsCancelMulZero α := - { mul_left_cancel_of_ne_zero := fun ha h ↦ by - rw [← sub_eq_zero, ← mul_sub] at h - exact sub_eq_zero.1 ((eq_zero_or_eq_zero_of_mul_eq_zero h).resolve_left ha) - mul_right_cancel_of_ne_zero := fun hb h ↦ by - rw [← sub_eq_zero, ← sub_mul] at h - exact sub_eq_zero.1 ((eq_zero_or_eq_zero_of_mul_eq_zero h).resolve_right hb) } + IsCancelMulZero α where + mul_left_cancel_of_ne_zero ha h := by + rw [← sub_eq_zero, ← mul_sub] at h + exact sub_eq_zero.1 ((eq_zero_or_eq_zero_of_mul_eq_zero h).resolve_left ha) + mul_right_cancel_of_ne_zero hb h := by + rw [← sub_eq_zero, ← sub_mul] at h + exact sub_eq_zero.1 ((eq_zero_or_eq_zero_of_mul_eq_zero h).resolve_right hb) /-- In a ring, `IsCancelMulZero` and `NoZeroDivisors` are equivalent. -/ lemma isCancelMulZero_iff_noZeroDivisors [NonUnitalNonAssocRing α] : @@ -165,7 +153,7 @@ lemma NoZeroDivisors.to_isDomain [Ring α] [h : Nontrivial α] [NoZeroDivisors IsDomain α := { NoZeroDivisors.to_isCancelMulZero α, h with .. } -instance (priority := 100) IsDomain.to_noZeroDivisors [Ring α] [IsDomain α] : +instance (priority := 100) IsDomain.to_noZeroDivisors [Semiring α] [IsDomain α] : NoZeroDivisors α := IsRightCancelMulZero.to_noZeroDivisors α diff --git a/Mathlib/Algebra/Ring/Hom/Basic.lean b/Mathlib/Algebra/Ring/Hom/Basic.lean index 90d26a14ce0cb..10cf5b59b7a95 100644 --- a/Mathlib/Algebra/Ring/Hom/Basic.lean +++ b/Mathlib/Algebra/Ring/Hom/Basic.lean @@ -49,9 +49,7 @@ end Semiring end RingHom /-- Pullback `IsDomain` instance along an injective function. -/ -protected theorem Function.Injective.isDomain [Ring α] [IsDomain α] [Ring β] (f : β →+* α) - (hf : Injective f) : IsDomain β := by - haveI := pullback_nonzero f f.map_zero f.map_one - haveI := IsRightCancelMulZero.to_noZeroDivisors α - haveI := hf.noZeroDivisors f f.map_zero f.map_mul - exact NoZeroDivisors.to_isDomain β +protected theorem Function.Injective.isDomain [Semiring α] [IsDomain α] [Semiring β] {F} + [FunLike F β α] [MonoidWithZeroHomClass F β α] (f : F) (hf : Injective f) : IsDomain β where + __ := domain_nontrivial f (map_zero _) (map_one _) + __ := hf.isCancelMulZero f (map_zero _) (map_mul _) diff --git a/Mathlib/Data/Complex/Basic.lean b/Mathlib/Data/Complex/Basic.lean index e1d6851ed9e2f..c14f55bdce967 100644 --- a/Mathlib/Data/Complex/Basic.lean +++ b/Mathlib/Data/Complex/Basic.lean @@ -283,7 +283,7 @@ theorem equivRealProdAddHom_symm_apply (p : ℝ × ℝ) : diamond from the other actions they inherit through the `ℝ`-action on `ℂ` and action transitivity defined in `Data.Complex.Module`. -/ instance : Nontrivial ℂ := - pullback_nonzero re rfl rfl + domain_nontrivial re rfl rfl -- Porting note: moved from `Module/Data/Complex/Basic.lean` namespace SMul diff --git a/Mathlib/Logic/Equiv/TransferInstance.lean b/Mathlib/Logic/Equiv/TransferInstance.lean index 2b78ff6212187..66dfeace0399c 100644 --- a/Mathlib/Logic/Equiv/TransferInstance.lean +++ b/Mathlib/Logic/Equiv/TransferInstance.lean @@ -186,7 +186,7 @@ theorem ringEquiv_symm_apply (e : α ≃ β) [Add β] [Mul β] (b : β) : by variable (α) in /-- Shrink `α` to a smaller universe preserves ring structure. -/ -noncomputable def _root_.Shrink.ringEquiv [Small.{v} α] [Ring α] : Shrink.{v} α ≃+* α := +noncomputable def _root_.Shrink.ringEquiv [Small.{v} α] [Add α] [Mul α] : Shrink.{v} α ≃+* α := (equivShrink α).symm.ringEquiv /-- Transfer `Semigroup` across an `Equiv` -/ @@ -471,10 +471,10 @@ noncomputable instance [Small.{v} α] [Nontrivial α] : Nontrivial (Shrink.{v} (equivShrink α).symm.nontrivial /-- Transfer `IsDomain` across an `Equiv` -/ -protected theorem isDomain [Ring α] [Ring β] [IsDomain β] (e : α ≃+* β) : IsDomain α := +protected theorem isDomain [Semiring α] [Semiring β] [IsDomain β] (e : α ≃+* β) : IsDomain α := Function.Injective.isDomain e.toRingHom e.injective -noncomputable instance [Small.{v} α] [Ring α] [IsDomain α] : IsDomain (Shrink.{v} α) := +noncomputable instance [Small.{v} α] [Semiring α] [IsDomain α] : IsDomain (Shrink.{v} α) := Equiv.isDomain (Shrink.ringEquiv α) /-- Transfer `NNRatCast` across an `Equiv` -/ diff --git a/Mathlib/RingTheory/Ideal/Prime.lean b/Mathlib/RingTheory/Ideal/Prime.lean index 321e1b9b104eb..afb9e63f85dcf 100644 --- a/Mathlib/RingTheory/Ideal/Prime.lean +++ b/Mathlib/RingTheory/Ideal/Prime.lean @@ -3,7 +3,6 @@ Copyright (c) 2018 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Chris Hughes, Mario Carneiro -/ -import Mathlib.Algebra.Ring.Regular import Mathlib.RingTheory.Ideal.Lattice /-! diff --git a/Mathlib/RingTheory/LocalRing/RingHom/Basic.lean b/Mathlib/RingTheory/LocalRing/RingHom/Basic.lean index fa149868d151e..0546a5cd3fb8e 100644 --- a/Mathlib/RingTheory/LocalRing/RingHom/Basic.lean +++ b/Mathlib/RingTheory/LocalRing/RingHom/Basic.lean @@ -55,7 +55,7 @@ alias isLocalRingHom_of_comp := isLocalHom_of_comp /-- If `f : R →+* S` is a local ring hom, then `R` is a local ring if `S` is. -/ theorem RingHom.domain_isLocalRing {R S : Type*} [CommSemiring R] [CommSemiring S] [IsLocalRing S] (f : R →+* S) [IsLocalHom f] : IsLocalRing R := by - haveI : Nontrivial R := pullback_nonzero f f.map_zero f.map_one + haveI : Nontrivial R := f.domain_nontrivial apply IsLocalRing.of_nonunits_add intro a b simp_rw [← map_mem_nonunits_iff f, f.map_add] diff --git a/Mathlib/RingTheory/Polynomial/IntegralNormalization.lean b/Mathlib/RingTheory/Polynomial/IntegralNormalization.lean index 967c8dbcef67d..ee4f04ec120b8 100644 --- a/Mathlib/RingTheory/Polynomial/IntegralNormalization.lean +++ b/Mathlib/RingTheory/Polynomial/IntegralNormalization.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Johannes Hölzl, Kim Morrison, Jens Wagemaker, Andrew Yang, Yuyang Zhao -/ import Mathlib.Algebra.Polynomial.Monic -import Mathlib.Algebra.Ring.Regular import Mathlib.RingTheory.Polynomial.ScaleRoots /-! From 93a38c113b9d76645c526dca1c5644d702f67a58 Mon Sep 17 00:00:00 2001 From: grunweg Date: Mon, 9 Dec 2024 09:39:53 +0000 Subject: [PATCH 586/829] feat: add a file tracking results from the 1000+ theorems project (#19634) Add a new file `1000.yaml`, tracking the formalisation status of theorems from the [1000+ theorems project](https://1000-plus.github.io). Make `check-yaml` handle this file also. Future PRs will - expose this information on the webpage - add infrastructure for synchronising this file with the upstream repository - could add an attribute similar to the `stacks` attribute, and generate this file from the tagging - add information on existing formalisation projects in Lean. --- .github/build.in.yml | 2 +- .github/workflows/bors.yml | 2 +- .github/workflows/build.yml | 2 +- .github/workflows/build_fork.yml | 2 +- docs/1000.yaml | 3612 ++++++++++++++++++++++++++++++ scripts/check-yaml.lean | 4 +- scripts/yaml_check.py | 25 +- 7 files changed, 3638 insertions(+), 11 deletions(-) create mode 100644 docs/1000.yaml diff --git a/.github/build.in.yml b/.github/build.in.yml index 6a1d17a8cf495..c56faf4061fea 100644 --- a/.github/build.in.yml +++ b/.github/build.in.yml @@ -240,7 +240,7 @@ jobs: - name: check declarations in db files run: | - python3 scripts/yaml_check.py docs/100.yaml docs/overview.yaml docs/undergrad.yaml + python3 scripts/yaml_check.py docs/100.yaml docs/1000.yaml docs/overview.yaml docs/undergrad.yaml lake exe check-yaml - name: generate our import graph diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index 87e3093f9b131..869b4d379c959 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -250,7 +250,7 @@ jobs: - name: check declarations in db files run: | - python3 scripts/yaml_check.py docs/100.yaml docs/overview.yaml docs/undergrad.yaml + python3 scripts/yaml_check.py docs/100.yaml docs/1000.yaml docs/overview.yaml docs/undergrad.yaml lake exe check-yaml - name: generate our import graph diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 5ae1b33476e94..25c242698e644 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -257,7 +257,7 @@ jobs: - name: check declarations in db files run: | - python3 scripts/yaml_check.py docs/100.yaml docs/overview.yaml docs/undergrad.yaml + python3 scripts/yaml_check.py docs/100.yaml docs/1000.yaml docs/overview.yaml docs/undergrad.yaml lake exe check-yaml - name: generate our import graph diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index 79bfe1d0f3d0a..9848f1817b72d 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -254,7 +254,7 @@ jobs: - name: check declarations in db files run: | - python3 scripts/yaml_check.py docs/100.yaml docs/overview.yaml docs/undergrad.yaml + python3 scripts/yaml_check.py docs/100.yaml docs/1000.yaml docs/overview.yaml docs/undergrad.yaml lake exe check-yaml - name: generate our import graph diff --git a/docs/1000.yaml b/docs/1000.yaml new file mode 100644 index 0000000000000..8126486a00af9 --- /dev/null +++ b/docs/1000.yaml @@ -0,0 +1,3612 @@ +# This file tracks the formalisation status of theorems from the experimental +# 1000+ theorems project (https://1000-plus.github.io/). +# This file and associated infrastructure are work in progress. +# +# Current TODOs/unresolved questions: +# - should this file include further metadata, such as repository links, author names or dates? +# - add infrastructure for updating the information upstream when formalisations are added +# - perhaps add a wikidata version of the stacks tag, and generate this file automatically +# (can this handle formalisation of *statements* also?) +# - complete the information which theorems are already formalised + +Q11518: + title: Pythagorean theorem + +Q12524: + title: Schwenk's theorem + +Q26708: + title: Binomial theorem + +Q32182: + title: Balinski's theorem + +Q33481: + title: Arrow's impossibility theorem + +Q98831: + title: Multiplication theorem + +Q117553: + title: Modigliani–Miller theorem + +Q119450: + title: Brianchon's theorem + +Q132427: + title: Rademacher's theorem + +Q132469: + title: Fermat's Last Theorem + +Q137164: + title: Besicovitch covering theorem + +Q154210: + title: CPCTC + +Q164262: + title: Lebesgue covering dimension + +Q172298: + title: Lasker–Noether theorem + +Q174955: + title: Mihăilescu's theorem + +Q179208: + title: Cayley's theorem + +Q179467: + title: Fourier theorem + +Q179692: + title: Independence of the axiom of choice + +Q180345X: + title: Integral root theorem + +Q180345: + title: Rational root theorem + +Q182505: + title: Bayes' theorem + +Q184410: + title: Four color theorem + +Q188295: + title: Fermat's little theorem + +Q188745: + title: Classification of Platonic solids + +Q189136: + title: Mean value theorem + +Q190026: + title: Ford's theorem + +Q190391X: + title: Lyapunov's central limit theorem + +Q190391: + title: Central limit theorem + +Q190556: + title: De Moivre's theorem + +Q191693: + title: Lebesgue's decomposition theorem + +Q192760: + title: Fundamental theorem of algebra + +Q193286: + title: Rolle's theorem + +Q193878: + title: Chinese remainder theorem + +Q193882: + title: Menelaus's theorem + +Q193910: + title: Euler's theorem + +Q194919: + title: Inverse eigenvalues theorem + +Q195133: + title: Wigner–Eckart theorem + +Q200787: + title: Gödel's incompleteness theorem + +Q203565: + title: Solutions of a general cubic equation + +Q204884: + title: Löb's theorem + +Q205966: + title: Critical line theorem + +Q207014: + title: Independence of the parallel postulate + +Q207244: + title: Sela's theorem + +Q208416: + title: Independence of the continuum hypothesis + url: https://github.com/flypitch/flypitch + identifiers: + - independence_of_CH + author: Floris van Doorn and Jesse Michael Han + date: 2019-09-17 + +Q208756: + title: Castelnuovo theorem + +Q211981: + title: Hartman–Grobman theorem + +Q213603: + title: Ceva's theorem + +Q220680: + title: Banach fixed-point theorem + +Q225973: + title: Ore's theorem + +Q226014: + title: Poincaré recurrence theorem + +Q230848: + title: Lamé’s theorem + +Q240950: + title: Faltings's theorem + +Q241868: + title: Thévenin's theorem + +Q242045: + title: Monotone class theorem + +Q245098: + title: Intermediate value theorem + +Q245098X: + title: Bolzano's theorem + +Q245902: + title: Riesz–Thorin theorem + +Q248931: + title: Poincaré–Bendixson theorem + +Q253214: + title: Heine–Borel theorem + +Q255996: + title: Equipartition theorem + +Q256303: + title: Lax–Richtmyer theorem + +Q257387: + title: Vitali theorem + +Q258374: + title: Carathéodory's theorem + +Q260928: + title: Jordan curve theorem + +Q266291: + title: Bloch's theorem + +Q268031: + title: Abel's binomial theorem + +Q268132: + title: Gauss–Wantzel theorem + +Q273037: + title: Bondy–Chvátal theorem + +Q276082: + title: Wilson's theorem + +Q280116: + title: Odd number theorem + +Q282331: + title: Addition theorem + +Q282649: + title: Pentagonal number theorem + +Q283690: + title: Arrival theorem + +Q284960: + title: Mordell–Weil theorem + +Q285719: + title: Thales's theorem + +Q287347: + title: Gradient theorem + +Q288465: + title: Orbit-stabilizer theorem + +Q303402: + title: Rank–nullity theorem + +Q318767: + title: Abel's theorem + +Q321237: + title: Green's theorem + +Q332465: + title: Absolute convergence theorem + +Q338886: + title: Divergence theorem + +Q339495: + title: Symphonic theorem + +Q356895: + title: Adiabatic theorem + +Q357192: + title: Holland's schema theorem + +Q357858: + title: Freyd's adjoint functor theorem + +Q371685: + title: Kraft–McMillan theorem + +Q372037: + title: Seifert–van Kampen theorem + +Q376166: + title: Cut-elimination theorem + +Q377047: + title: Butterfly theorem + +Q377276: + title: Cook's theorem + +Q379048: + title: Riemann–Roch theorem + +Q380576: + title: Measurable Riemann mapping theorem + +Q383238: + title: Hartogs's extension theorem + +Q384142: + title: Herbrand–Ribet theorem + +Q386292: + title: Prime number theorem + +Q388525: + title: Bell's theorem + +Q401319: + title: Puiseux's theorem + +Q401905: + title: Skorokhod's representation theorem + +Q420714: + title: Akra–Bazzi theorem + +Q422187: + title: Myhill–Nerode theorem + +Q425432: + title: Laurent expansion theorem + +Q428134: + title: Gauss–Markov theorem + +Q459547: + title: Ptolemy's theorem + +Q467205: + title: Fundamental theorems of welfare economics + +Q467756: + title: Stokes's theorem + +Q468391: + title: Bolzano–Weierstrass theorem + +Q470877: + title: Stirling's theorem + +Q472883: + title: Quadratic reciprocity theorem + +Q474881: + title: Cantor's theorem + +Q476776: + title: Solutions of a general quartic equation + +Q487132: + title: Infinite monkey theorem + +Q488541: + title: Fisher separation theorem + +Q495412: + title: Jordan–Schönflies theorem + +Q504843: + title: Mycielski's theorem + +Q505798: + title: Lagrange's theorem + +Q510197: + title: Tutte theorem + +Q512897: + title: Brooks's theorem + +Q523832: + title: Carathéodory–Jacobi–Lie theorem + +Q523876: + title: Spin–statistics theorem + +Q528955: + title: Kochen–Specker theorem + +Q528987: + title: Von Neumann bicommutant theorem + +Q530152: + title: Picard–Lindelöf theorem + +Q535366: + title: Montel's theorem + +Q536640: + title: Hall's marriage theorem + +Q537618: + title: Banach–Alaoglu theorem + +Q544292: + title: Lagrange inversion theorem + +Q544369: + title: Glivenko–Cantelli theorem + +Q550402: + title: Dirichlet's theorem on arithmetic progressions + +Q552367: + title: Berge's theorem + +Q567843: + title: Hahn–Mazurkiewicz theorem + +Q570779: + title: Vieta's formulas + +Q574902: + title: Tarski's indefinability theorem + +Q576478: + title: Liouville's theorem + +Q578555: + title: Noether's theorem + +Q579515: + title: Rao–Blackwell theorem + +Q583147: + title: Sard's theorem + +Q586051: + title: Haag–Łopuszański–Sohnius theorem + +Q587081: + title: Kolmogorov–Arnold–Moser theorem + +Q588218: + title: Koebe 1/4 theorem + +Q594571: + title: Van Aubel's theorem + +Q595466: + title: Dini's theorem + +Q599876: + title: Cochran's theorem + +Q608294: + title: Max flow min cut theorem + +Q609612: + title: Knaster–Tarski theorem + +Q612021: + title: Gibbard–Satterthwaite theorem + +Q617417: + title: Isoperimetric theorem + +Q619985: + title: Multinomial theorem + +Q620602: + title: Virial theorem + +Q631561: + title: Closed range theorem + +Q632546X: + title: Sylvester's theorem + +Q632546: + title: Bertrand's postulate + +Q637418: + title: Napoleon's theorem + +Q642620: + title: Krein–Milman theorem + +Q643513: + title: Riesz–Fischer theorem + +Q643826: + title: Slutsky's theorem + +Q646523: + title: Pick's theorem + +Q649469: + title: Modularity theorem + +Q649977: + title: Shirshov–Cohn theorem + +Q650738: + title: Folk theorem + +Q651593: + title: Norton's theorem + +Q656176: + title: Schauder fixed-point theorem + +Q656198: + title: Maschke's theorem + +Q656645: + title: Hilbert's basis theorem + +Q656772: + title: Cayley–Hamilton theorem + +Q657469: + title: Lefschetz fixed-point theorem + +Q657469X: + title: Lefschetz–Hopf theorem + +Q657482: + title: Abel–Ruffini theorem + +Q657903: + title: Hilbert–Waring theorem + +Q660799: + title: Darboux's theorem + +Q670235: + title: Fundamental theorem of arithmetic + +Q671663: + title: Coleman–Mandula theorem + +Q676413: + title: Dandelin's theorem + +Q679800: + title: Nyquist–Shannon sampling theorem + +Q681406: + title: Euler's rotation theorem + +Q685437: + title: Titchmarsh theorem + +Q693083: + title: Soundness theorem + +Q716171: + title: Erdős–Ginzburg–Ziv theorem + +Q718875: + title: Erdős–Ko–Rado theorem + +Q719966: + title: Toda's theorem + +Q720469: + title: Chevalley–Warning theorem + +Q721695: + title: Szemerédi–Trotter theorem + +Q727102: + title: Fermat's theorem (stationary points) + +Q729359: + title: Hadamard three-circle theorem + +Q730222: + title: Ham sandwich theorem + +Q733081: + title: Impossibility of angle trisection + +Q737851: + title: Banach–Tarski theorem + +Q737892: + title: Bendixson–Dulac theorem + +Q739403: + title: Stewart's theorem + +Q740093: + title: Peano existence theorem + +Q742833: + title: Gauss–Bonnet theorem + +Q744440: + title: Immerman–Szelepcsényi theorem + +Q748233: + title: Sylvester–Gallai theorem + +Q751120: + title: Thue–Siegel–Roth theorem + +Q752375: + title: Extreme value theorem + +Q755986: + title: Atiyah–Bott fixed-point theorem + +Q755991: + title: Atiyah–Singer index theorem + +Q756946: + title: Lagrange's four-square theorem + +Q764287: + title: Van der Waerden's theorem + +Q765987: + title: Heine–Cantor theorem + +Q766722: + title: Liouville's theorem + +Q776578: + title: Artin–Wedderburn theorem + +Q777924: + title: Tijdeman's theorem + +Q779220: + title: Hilbert's syzygy theorem + +Q780763: + title: 15 and 290 theorems + +Q784049: + title: Blaschke selection theorem + +Q790236: + title: Baranyai's theorem + +Q791258: + title: Basu's theorem + +Q791663: + title: Beatty's theorem + +Q794269: + title: Betti's theorem + +Q810431: + title: Basel problem + +Q828284: + title: Parallel axis theorem + +Q830124: + title: Girsanov's theorem + +Q830513: + title: Residue theorem + +Q834025: + title: Cauchy integral theorem + +Q834211: + title: Wallace–Bolyai–Gerwien theorem + +Q837506: + title: Theorem on friends and strangers + +Q837551: + title: Frobenius theorem + +Q841893: + title: Desargues's theorem + +Q842953: + title: Lebesgue's density theorem + +Q844612: + title: Brun's theorem + +Q845088: + title: Pappus's centroid theorem + +Q846840: + title: Erdős–Kac theorem + +Q848092: + title: Borsuk–Ulam theorem + +Q848375: + title: Implicit function theorem + +Q848810: + title: Kronecker's theorem + +Q850495: + title: Wick's theorem + +Q851166: + title: Pappus's hexagon theorem + +Q852183: + title: Viviani's theorem + +Q852973: + title: Euler's polyhedron theorem + +Q853067: + title: Solutions to Pell's equation + +Q856032: + title: CAP theorem + +Q856158: + title: Midy's theorem + +Q857089: + title: de Branges's theorem + +Q858978: + title: Castigliano's first and second theorems + +Q859122: + title: Cauchy–Hadamard theorem + +Q865665: + title: Birkhoff's theorem + +Q866116: + title: Hahn–Banach theorem + +Q867141: + title: Ehresmann's theorem + +Q867839: + title: Rosser's theorem + +Q869647: + title: Szpilrajn extension theorem + +Q872088: + title: Boolean prime ideal theorem + +Q877489: + title: Apollonius's theorem + +Q890875: + title: Bohr–van Leeuwen theorem + +Q897769: + title: König's theorem + +Q899002: + title: Pascal's theorem + +Q899853: + title: H-theorem + +Q900831: + title: Fluctuation theorem + +Q902052: + title: Gödel's completeness theorem + +Q902618: + title: Floquet's theorem + +Q906809: + title: Hellmann–Feynman theorem + +Q909517: + title: Feit–Thompson theorem + +Q913447: + title: Morley's trisector theorem + +Q913849: + title: Minimax theorem + +Q914517: + title: Fermat's theorem on sums of two squares + +Q915474: + title: Robin's theorem + +Q918099: + title: Ramsey's theorem + +Q922012: + title: Green–Tao theorem + +Q922367: + title: Master theorem (analysis of algorithms) + +Q925224: + title: Sonnenschein–Mantel–Debreu Theorem + +Q925854: + title: Angle bisector theorem + +Q927051: + title: Riemann mapping theorem + +Q928813: + title: Menger's theorem + +Q929547: + title: Myers theorem + +Q931001: + title: Inverse function theorem + +Q931001X: + title: Constant rank theorem + +Q931404: + title: Bertrand–Diquet–Puiseux theorem + +Q939927: + title: Stone–Weierstrass theorem + decl: ContinuousMap.subalgebra_topologicalClosure_eq_top_of_separatesPoints + author: Scott Morrison and Heather Macbeth + date: 2021-05-01 + +Q942046: + title: Schwarz–Ahlfors–Pick theorem + +Q943246: + title: Wedderburn's little theorem + +Q944297: + title: Open mapping theorem + +Q948664: + title: Kneser's theorem + +Q951327: + title: Pasch's theorem + +Q953062: + title: Reynolds transport theorem + +Q954210: + title: Savitch's theorem + +Q965294: + title: Thabit ibn Qurra's theorem + +Q966837: + title: Cramér–Wold theorem + +Q967972: + title: Open mapping theorem + +Q973359: + title: Rado's theorem + +Q974405: + title: Hille–Yosida theorem + +Q976033: + title: Gelfond–Schneider theorem + +Q976607: + title: Erdős–Szekeres theorem + +Q977912: + title: Stolper–Samuelson theorem + +Q978688: + title: Gershgorin circle theorem + +Q985009: + title: Helmholtz's theorems + +Q994401: + title: Four-vertex theorem + +Q995926: + title: Perpendicular axis theorem + +Q999783: + title: Buckingham π theorem + +Q1008566: + title: Gauss–Lucas theorem + +Q1032886: + title: Hales–Jewett theorem + +Q1033910: + title: Cantor–Bernstein–Schroeder theorem + +Q1037559: + title: Pólya enumeration theorem + +Q1038716: + title: Identity theorem + +Q1046232: + title: Szemerédi's theorem + +Q1047749: + title: Turán's theorem + +Q1048589: + title: Hohenberg–Kohn theorems + +Q1048874: + title: Gauss's Theorema Egregium + +Q1050037: + title: Coase theorem + +Q1050203: + title: Cantor's intersection theorem + +Q1050932: + title: Hartogs's theorem + +Q1051404: + title: Cesàro's theorem + +Q1052021: + title: Craig's interpolation theorem + +Q1052678: + title: Baire category theorem + +Q1052752: + title: Stolz–Cesàro theorem + +Q1054106: + title: Cox's theorem + +Q1056283: + title: May's theorem + +Q1057919: + title: Sylow theorems + +Q1058662: + title: Darboux's theorem + +Q1059151: + title: FWL theorem + +Q1064471: + title: Earnshaw's theorem + +Q1065257: + title: Squeeze theorem + +Q1065966: + title: Isomorphism theorem + +Q1067156: + title: Dominated convergence theorem + +Q1067156X: + title: Bounded convergence theorem + +Q1068085: + title: Closed graph theorem + +Q1068283: + title: Löwenheim–Skolem theorem + +Q1068385: + title: Clairaut's theorem + +Q1068976: + title: Hilbert's Nullstellensatz + +Q1075398: + title: Jung's theorem + +Q1076274: + title: Choquet–Bishop–de Leeuw theorem + +Q1077462: + title: König's theorem + +Q1077741: + title: Hairy ball theorem + +Q1082910: + title: Euler's partition theorem + +Q1095330: + title: Apéry's theorem + +Q1097021: + title: Minkowski's theorem + +Q1103054: + title: Lions–Lax–Milgram theorem + +Q1119094: + title: Paley–Wiener theorem + +Q1121027: + title: Krylov–Bogolyubov theorem + +Q1125462: + title: Dinostratus' theorem + +Q1129636: + title: Lehmann–Scheffé theorem + +Q1130846: + title: Ladner's theorem + +Q1131601: + title: Sklar's theorem + +Q1132952: + title: Euler's theorem on homogeneous functions + +Q1134296: + title: Supporting hyperplane theorem + +Q1134776: + title: Dilworth's theorem + +Q1135706: + title: Harnack's theorem + +Q1136043: + title: Hurwitz's theorem + +Q1136262: + title: Optical theorem + +Q1137014: + title: Tychonoff's theorem + +Q1137206: + title: Taylor's theorem + +Q1139041: + title: Cauchy's theorem + +Q1139430: + title: Hurwitz's theorem + +Q1139524: + title: Rationality theorem + +Q1140119: + title: Morera's theorem + +Q1140200: + title: PCP theorem + +Q1141747: + title: Carnot's theorem + +Q1143540: + title: Heckscher–Ohlin theorem + +Q1144897: + title: Brouwer fixed-point theorem + +Q1146791: + title: Fermat polygonal number theorem + +Q1148215: + title: Mitchell's embedding theorem + +Q1149022: + title: Fubini's theorem + +Q1149185X: + title: Kirby–Paris theorem + +Q1149185: + title: Goodstein's theorem + +Q1149458: + title: Compactness theorem + +Q1149522: + title: Lami's theorem + +Q1152521: + title: Mohr–Mascheroni theorem + +Q1153441: + title: Goldstone's theorem + +Q1153584: + title: Monotone convergence theorem + +Q1166774: + title: Stone's representation theorem for Boolean algebras + +Q1177508: + title: Rédei's theorem + +Q1179446: + title: De Rham's theorem + +Q1182249: + title: Deduction theorem + +Q1186134: + title: Poncelet–Steiner theorem + +Q1186808: + title: Lucas's theorem + +Q1187646: + title: Fundamental theorem on homomorphisms + +Q1188048: + title: Barbier's theorem + +Q1188392: + title: Zeckendorf's theorem + +Q1188504: + title: Pitman–Koopman–Darmois theorem + +Q1191319: + title: Radon–Nikodym theorem + +Q1191709: + title: Egorov's theorem + +Q1191862: + title: Rouché's theorem + +Q1194053: + title: Metrization theorems + +Q1196538: + title: Descartes's theorem + +Q1196729: + title: Mertens's theorems + +Q1202608: + title: De Gua's theorem + +Q1209546: + title: Kaplansky density theorem + +Q1217677: + title: Fundamental theorem of calculus + +Q1227061: + title: Khinchin's theorem + +Q1227702: + title: Dirichlet's unit theorem + +Q1227703: + title: Dirichlet's approximation theorem + +Q1242398: + title: Doob decomposition theorem + +Q1243340: + title: Birkhoff–Von Neumann theorem + +Q1249069: + title: Richardson's theorem + +Q1259814: + title: Nielsen–Schreier theorem + +Q1276318: + title: Schwartz kernel theorem + +Q1305766: + title: Poincaré–Hopf theorem + +Q1306092: + title: Nash embedding theorem + +Q1306095: + title: Whitney embedding theorem + +Q1307676: + title: Künneth theorem + +Q1308502: + title: Church–Rosser theorem + +Q1315949: + title: Mittag-Leffler's theorem + +Q1317350: + title: Chen's theorem + +Q1317367: + title: Japanese theorem for concyclic polygons + +Q1322892: + title: Dirac's theorems + +Q1330788: + title: Weierstrass factorization theorem + +Q1340623: + title: Classification of finite simple groups + +Q1346677: + title: Tietze extension theorem + +Q1347094: + title: Alternate Interior Angles Theorem + +Q1348736: + title: Rybczynski theorem + +Q1349282: + title: Eilenberg–Zilber theorem + +Q1357684: + title: Riesz representation theorem + +Q1361031: + title: Frobenius theorem + +Q1361393: + title: Luzin's theorem + +Q1366581: + title: Erdős–Rado theorem + +Q1369453: + title: Kronecker–Weber theorem + +Q1370316: + title: Casey's theorem + +Q1376515: + title: No-hair theorem + +Q1396592: + title: Franel–Landau theorem + +Q1420905: + title: Hilbert's theorem 90 + +Q1422083: + title: Ryll-Nardzewski fixed-point theorem + +Q1423818: + title: Euler's theorem in geometry + +Q1424481: + title: Frucht's theorem + +Q1425077: + title: Spectral theorem + +Q1425308: + title: Brahmagupta theorem + +Q1425529: + title: Chebotarev's density theorem + +Q1426292: + title: Banach–Steinhaus theorem + +Q1434158: + title: Fluctuation dissipation theorem + +Q1434805: + title: Max Noether's theorem + +Q1443036: + title: Parseval's theorem + +Q1452678: + title: Penrose–Hawking singularity theorems + +Q1455794: + title: Freudenthal suspension theorem + +Q1457052: + title: Well-ordering theorem + +Q1471282: + title: Radon's theorem + +Q1472120: + title: Hellinger–Toeplitz theorem + +Q1477053: + title: Arzelà–Ascoli theorem + +Q1505529: + title: Runge's theorem + +Q1506253: + title: Euclid's theorem + +Q1529941: + title: Peter–Weyl theorem + +Q1530275: + title: Dunford–Pettis theorem + +Q1535225: + title: Łoś' theorem + +Q1542114: + title: Bézout's theorem + +Q1543149: + title: Sokhatsky–Weierstrass theorem + +Q1547975: + title: Mermin–Wagner theorem + +Q1551745: + title: Binomial inverse theorem + +Q1555342: + title: Reeh–Schlieder theorem + +Q1564541: + title: Perron–Frobenius theorem + +Q1566341: + title: Hindman's theorem + +Q1566372: + title: Haag's theorem + +Q1568612: + title: Marden's theorem + +Q1568811: + title: Hahn decomposition theorem + +Q1572474: + title: Lindemann–Weierstrass theorem + +Q1576235: + title: Lochs's theorem + +Q1602819: + title: Lumer–Phillips theorem + +Q1614464: + title: Cauchy–Kowalevski theorem + +Q1621180: + title: Orlicz–Pettis theorem + +Q1630588: + title: Hurwitz's theorem + +Q1631749: + title: Lester's theorem + +Q1631992: + title: Steiner–Lehmus theorem + +Q1632301: + title: Sturm's theorem + +Q1632433: + title: Helly's theorem + +Q1637085: + title: Poynting's theorem + +Q1660147: + title: Mazur–Ulam theorem + +Q1668861: + title: Picard theorem + +Q1683356: + title: Japanese theorem for concyclic quadrilaterals + +Q1687147: + title: Sprague–Grundy theorem + +Q1694565: + title: Vinogradov's theorem + +Q1699024: + title: John ellipsoid + +Q1704878: + title: Appell–Humbert theorem + +Q1716166: + title: Mercer's theorem + +Q1720410: + title: Hjelmslev's theorem + +Q1724049: + title: Wolstenholme's theorem + +Q1727461: + title: Intersecting secants theorem + +Q1739994: + title: Brauer–Suzuki theorem + +Q1751105: + title: Blum's speedup theorem + +Q1751823: + title: No-cloning theorem + +Q1752516: + title: Riemann series theorem + +Q1752621: + title: Sylvester's law of inertia + +Q1752988: + title: Kutta–Joukowski theorem + +Q1756194: + title: Shannon–Hartley theorem + +Q1765521: + title: Doob's martingale convergence theorems + +Q1766814: + title: Smn theorem + +Q1785610: + title: Poncelet's closure theorem + +Q1785872: + title: Berry–Esséen theorem + +Q1786292: + title: Schur's theorem + +Q1786419: + title: Kramers' theorem + +Q1789863: + title: Routh's theorem + +Q1809539: + title: Pizza theorem + +Q1816952: + title: Schur's lemma + +Q1841237: + title: Fubini's theorem on differentiation + +Q1855610: + title: Theorem of de Moivre–Laplace + +Q1888583: + title: Holditch's theorem + +Q1893717: + title: Rice's theorem + +Q1896455: + title: Varignon's theorem + +Q1899432: + title: Grothendieck–Hirzebruch–Riemann–Roch theorem + +Q1914905: + title: Ramanujan–Skolem's theorem + +Q1930577: + title: Herbrand's theorem + +Q1933521: + title: Kleene's recursion theorem + +Q1938251: + title: Hasse–Minkowski theorem + +Q1946642X: + title: S–cobordism theorem + +Q1946642: + title: H-cobordism theorem + +Q1964537: + title: Mergelyan's theorem + +Q1966978: + title: Lévy continuity theorem + +Q1974087: + title: Riemann's theorem on removable singularities + +Q1995468: + title: Hölder's theorem + +Q1996305: + title: Burnside's theorem + +Q2000090: + title: Art gallery theorem + +Q2008549: + title: Hilbert's theorem + +Q2013213: + title: Reuschle's theorem + +Q2022775: + title: Hilbert–Schmidt theorem + +Q2027347: + title: Optional stopping theorem + +Q2028341: + title: Ado's theorem + +Q2046647: + title: Karhunen–Loève theorem + +Q2063099: + title: Poincaré duality theorem + +Q2068227: + title: Artin–Schreier theorem + +Q2071632: + title: Rouché–Capelli theorem + +Q2093886: + title: Primitive element theorem + +Q2094241: + title: Hopf–Rinow theorem + +Q2096184: + title: Plancherel theorem + +Q2109761: + title: Uniformization theorem + +Q2120067: + title: Diller–Dress theorem + +Q2190744: + title: Feuerbach's theorem + +Q2197859: + title: Nicomachus's theorem + +Q2222575: + title: Shell theorem + +Q2226602: + title: Bauer–Fike theorem + +Q2226606: + title: Beckman–Quarles theorem + +Q2226610: + title: Banach–Mazur theorem + +Q2226624: + title: Bruck–Chowla–Ryser theorem + +Q2226625: + title: Commandino's theorem + +Q2226640: + title: Cramér’s decomposition theorem + +Q2226644: + title: Easton's theorem + +Q2226650: + title: Eberlein–Šmulian theorem + +Q2226677: + title: Gelfand–Mazur theorem + +Q2226682: + title: Gelfand–Naimark theorem + +Q2226691: + title: Kirchhoff's theorem + +Q2226695: + title: Hurwitz's automorphisms theorem + +Q2226698: + title: Kantorovich theorem + +Q2226709: + title: Krull–Schmidt theorem + +Q2226750: + title: Malgrange–Ehrenpreis theorem + +Q2226774: + title: König's theorem + +Q2226786: + title: Sperner's theorem + +Q2226800: + title: Schur–Zassenhaus theorem + +Q2226803: + title: Tennenbaum's theorem + +Q2226807: + title: Vantieghems theorem + +Q2226855: + title: Sarkovskii's theorem + +Q2226868: + title: Geometric mean theorem + +Q2226962: + title: Weierstrass–Casorati theorem + +Q2253746: + title: Bertrand's ballot theorem + +Q2266397: + title: Intersecting chords theorem + +Q2267175: + title: Tangent-secant theorem + +Q2269888: + title: Cohn's irreducibility criterion + +Q2270905: + title: Poincaré–Birkhoff–Witt theorem + +Q2275191: + title: Lebesgue differentiation theorem + +Q2277040: + title: Shannon's expansion theorem + +Q2287393: + title: Caristi fixed-point theorem + +Q2309682: + title: Sphere theorem + +Q2309760: + title: Looman–Menchoff theorem + +Q2310718: + title: Pitot theorem + +Q2338929: + title: Exchange theorem + +Q2345282: + title: Shannon's theorem + +Q2347139: + title: Carnot's theorem + +Q2360531: + title: Jordan–Schur theorem + +Q2374733: + title: Skolem–Mahler–Lech theorem + +Q2376918: + title: Abelian and Tauberian theorems + +Q2378270: + title: Thue's theorem + +Q2379128: + title: Lindström's theorem + +Q2379132: + title: Going-up and going-down theorems + +Q2394548: + title: Hurewicz theorem + +Q2411312: + title: Shannon's source coding theorem + +Q2428364: + title: Clausius theorem + +Q2449984: + title: Miquel's theorem + +Q2449984X: + title: Pivot theorem + +Q2471737: + title: Carathéodory's theorem + +Q2473965: + title: Ugly duckling theorem + +Q2478371: + title: Envelope theorem + +Q2482753: + title: Mann's theorem + +Q2495664: + title: Universal coefficient theorem + +Q2500406: + title: Liénard's theorem + +Q2500408: + title: Thébault's theorem + +Q2518048: + title: Kodaira vanishing theorem + +Q2524990: + title: Jackson's theorem + +Q2525646: + title: Jordan–Hölder theorem + +Q2533936: + title: Descartes's theorem on total angular defect + +Q2540738: + title: Cartan–Dieudonné theorem + +Q2558669: + title: Artin–Zorn theorem + +Q2588432: + title: Wold's theorem + +Q2600653: + title: Taylor–Proudman theorem + +Q2607007: + title: Hinge theorem + +Q2621667: + title: Nagata–Smirnov metrization theorem + +Q2631152: + title: Carathéodory's extension theorem + +Q2635326: + title: Structured program theorem + +Q2638931: + title: Convolution theorem + +Q2703389: + title: Linnik's theorem + +Q2733672: + title: Price's theorem + +Q2734084: + title: Bohr–Mollerup theorem + +Q2793600: + title: Stark–Heegner theorem + +Q2799491: + title: Ringel–Youngs theorem + +Q2800071: + title: Perfect graph theorem + +Q2862403: + title: Bombieri's theorem + +Q2893695: + title: Unmixedness theorem + +Q2916568: + title: Gomory's theorem + +Q2919401: + title: Ostrowski's theorem + +Q2981012: + title: Monge's theorem + +Q2984415: + title: Hajnal–Szemerédi theorem + +Q2993026: + title: Ankeny–Artin–Chowla theorem + +Q2993308: + title: Cameron–Erdős theorem + +Q2993319: + title: Mirsky–Newman theorem + +Q2995691: + title: Newlander–Niremberg theorem + +Q3001796: + title: Skolem–Noether theorem + +Q3007384: + title: Post's theorem + +Q3075250: + title: Hardy–Littlewood maximal theorem + +Q3077634: + title: Sahlqvist correspondence theorem + +Q3088644: + title: Goldstine theorem + +Q3154003: + title: Le Cam's theorem + +Q3180727: + title: Perlis theorem + +Q3229335: + title: Borel–Carathéodory theorem + +Q3229352: + title: Vitali covering theorem + +Q3345678: + title: Moore–Aronszajn theorem + +Q3424029: + title: Chasles's theorems + +Q3443426: + title: Bondy's theorem + +Q3458641: + title: Szegő limit theorems + +Q3502015: + title: Hasse's theorem on elliptic curves + +Q3505260: + title: Cayley–Salmon theorem + +Q3526969: + title: AF+BG theorem + +Q3526976: + title: Ax–Kochen theorem + +Q3526982: + title: Erdős–Anning theorem + +Q3526990: + title: Euler's theorem + +Q3526993: + title: Takagi existence theorem + +Q3526996: + title: Kolmogorov extension theorem + +Q3526998: + title: Excision theorem + +Q3527004: + title: BEST theorem + +Q3527008: + title: Balian–Low theorem + +Q3527009: + title: Baker's theorem + +Q3527011: + title: Beck's theorem + +Q3527015: + title: Bernstein's theorem + +Q3527017: + title: Beurling–Lax theorem + +Q3527019: + title: Birch's theorem + +Q3527034: + title: Cauchy's theorem + +Q3527040: + title: Chomsky–Schützenberger representation theorem + +Q3527044: + title: Pappus's area theorem + +Q3527054: + title: De Bruijn–Erdős theorem (graph theory) + +Q3527056: + title: De Bruijn–Erdős theorem (incidence geometry) + +Q3527064: + title: Donsker's theorem + +Q3527067: + title: F. and M. Riesz theorem + +Q3527071: + title: Fáry–Milnor theorem + +Q3527079: + title: Freiman's theorem + +Q3527088: + title: Haboush's theorem + +Q3527091: + title: Gromov's theorem on groups of polynomial growth + +Q3527093: + title: Hilbert–Speiser theorem + +Q3527095: + title: Jacobi's four-square theorem + +Q3527100: + title: Kruskal's tree theorem + +Q3527102: + title: Kruskal–Katona theorem + +Q3527110: + title: Lax–Wendroff theorem + +Q3527113: + title: Lie–Kolchin theorem + +Q3527118: + title: Mahler's theorem + +Q3527125: + title: Monsky's theorem + +Q3527129: + title: Müntz–Szász theorem + +Q3527132: + title: Nagell–Lutz theorem + +Q3527155: + title: Robertson–Seymour theorem + +Q3527162: + title: Sophie Germain's theorem + +Q3527166: + title: Steinhaus theorem + +Q3527171: + title: Synge's theorem + +Q3527185: + title: Whitehead theorem + +Q3527189: + title: Lattice theorem + +Q3527196: + title: Principal ideal theorem + +Q3527205: + title: Dimension theorem for vector spaces + +Q3527209: + title: Hasse norm theorem + +Q3527214: + title: Non-squeezing theorem + +Q3527215: + title: Hilbert projection theorem + +Q3527217: + title: M. Riesz extension theorem + +Q3527219: + title: Weierstrass preparation theorem + +Q3527223: + title: Crystallographic restriction theorem + +Q3527226: + title: Linear speedup theorem + +Q3527230: + title: Von Staudt–Clausen theorem + +Q3527233: + title: Kolmogorov's three-series theorem + +Q3527241: + title: Krull's principal ideal theorem + +Q3527245: + title: Six exponentials theorem + +Q3527247: + title: Six circles theorem + +Q3527250: + title: Hadamard three-lines theorem + +Q3527263: + title: Kleene fixed-point theorem + +Q3527279: + title: Wiener's tauberian theorem + +Q3534036: + title: Pompeiu's theorem + +Q3613173: + title: Tits alternative + +Q3661294: + title: Kelvin's circulation theorem + +Q3699212: + title: Saint-Venant's theorem + +Q3711848: + title: Sobolev embedding theorem + +Q3757563: + title: Intercept theorem + +Q3771212: + title: Proth's theorem + +Q3889376: + title: Carmichael's theorem + +Q3893473: + title: Tameness theorem + +Q3915398: + title: Brunn–Minkowski theorem + +Q3983968: + title: Sturm–Picone comparison theorem + +Q3983972: + title: Simplicial approximation theorem + +Q3984011: + title: Duggan–Schwartz theorem + +Q3984017: + title: Helly's selection theorem + +Q3984020: + title: Holmström's theorem + +Q3984052: + title: Equidistribution theorem + +Q3984053: + title: Fourier inversion theorem + +Q3984056: + title: No-communication theorem + +Q3984063: + title: Mostow rigidity theorem + +Q3984069: + title: Fredholm's theorem + +Q4054114: + title: ATS theorem + +Q4116459: + title: Niven's theorem + +Q4171386: + title: Soul theorem + +Q4179283: + title: Positive energy theorem + +Q4272300: + title: Initial value theorem + +Q4272645: + title: Final value theorem + +Q4378868: + title: Phragmén–Lindelöf theorem + +Q4378889: + title: Carathéodory's theorem + +Q4408070: + title: De Finetti's theorem + +Q4454910: + title: Ostrowski–Hadamard gap theorem + +Q4454919: + title: Aumann's agreement theorem + +Q4454925: + title: Edge-of-the-wedge theorem + +Q4454954: + title: Kirszbraun theorem + +Q4454958: + title: Min-max theorem + +Q4454973: + title: Lindelöf's theorem + +Q4454976: + title: Liouville's theorem + +Q4454989: + title: Meusnier's theorem + +Q4454996: + title: Novikov's compact leaf theorem + +Q4455003: + title: Pomeranchuk's theorem + +Q4455015: + title: Routh–Hurwitz theorem + +Q4455016: + title: Reeb sphere theorem + +Q4455023: + title: Sazonov's theorem + +Q4455025: + title: No wandering domain theorem + +Q4455030: + title: Stone's theorem on one-parameter unitary groups + +Q4455031: + title: Fáry's theorem + +Q4455033: + title: Fatou's theorem + +Q4455037: + title: Hardy's theorem + +Q4455043: + title: Cayley–Bacharach theorem + +Q4455046: + title: Monodromy theorem + +Q4634017: + title: 2π theorem + +Q4663326: + title: Hirzebruch–Riemann–Roch theorem + +Q4666729: + title: Abel–Jacobi theorem + +Q4666729X: + title: Abel's curve theorem + +Q4667501: + title: Abhyankar–Moh theorem + +Q4677985: + title: Acyclic models theorem + +Q4695203: + title: Four functions theorem + +Q4700718: + title: Akhiezer's theorem + +Q4712316: + title: Albert–Brauer–Hasse–Noether theorem + +Q4713046: + title: Alchian–Allen theorem + +Q4724004B: + title: Riemann's existence theorem + +Q4724004A: + title: Chow's theorem + +Q4734002: + title: Gromov–Ruh theorem + +Q4734844: + title: Alperin–Brauer–Gorenstein theorem + +Q4736422: + title: Denjoy theorem + +Q4746948: + title: Amitsur–Levitzki theorem + +Q4751110: + title: Analyst's traveling salesman theorem + +Q4751118: + title: Analytic Fredholm theorem + +Q4753992: + title: Anderson's theorem + +Q4756099: + title: Andreotti–Frankel theorem + +Q4783822: + title: Arithmetic Riemann–Roch theorem + +Q4788680: + title: Area theorem (conformal mapping) + +Q4801169: + title: Artin approximation theorem + +Q4801183: + title: Artin–Verdier duality theorem + +Q4801563: + title: Artstein's theorem + +Q4815879: + title: Atiyah–Segal completion theorem + +Q4815899: + title: Atkinson's theorem + +Q4826848: + title: Autonomous convergence theorem + +Q4827308: + title: Auxiliary polynomial theorem + +Q4830725: + title: Ax–Grothendieck theorem + +Q4832965: + title: Aztec diamond theorem + +Q4838119: + title: Babuška–Lax–Milgram theorem + +Q4848497: + title: Baily–Borel theorem + +Q4853767: + title: Banach–Stone theorem + +Q4857506: + title: Bapat–Beg theorem + +Q4865980: + title: Barwise compactness theorem + +Q4866385: + title: Base change theorems + +Q4866385X: + title: Proper base change theorem + +Q4877965: + title: Beauville–Laszlo theorem + +Q4878586: + title: Beck's monadicity theorem + +Q4880958: + title: Behnke–Stein theorem + +Q4884789: + title: Beltrami's theorem + +Q4884950: + title: Belyi's theorem + +Q4891633: + title: Berger–Kazdan comparison theorem + +Q4894565: + title: Bernstein's theorem + +Q4914101: + title: Bing's recognition theorem + +Q4916483: + title: Birkhoff–Grothendieck theorem + +Q4918128: + title: Bishop–Cannings theorem + +Q4927458: + title: Blondel's theorem + +Q4937691: + title: Bogoliubov–Parasyuk theorem + +Q4941364: + title: Bondareva–Shapley theorem + +Q4942215: + title: Bonnet theorem + +Q4944906: + title: Borel determinacy theorem + +Q4944908: + title: Borel fixed-point theorem + +Q4944923: + title: Borel–Bott–Weil theorem + +Q4944923X: + title: Borel–Weil theorem + +Q4948983: + title: Bott periodicity theorem + +Q4949984: + title: Bounded inverse theorem + +Q4950096: + title: Bourbaki–Witt theorem + +Q4956438: + title: Branching theorem + +Q4958218: + title: Brauer's theorem on induced characters + +Q4958221: + title: Brauer's three main theorems + +Q4958227: + title: Brauer–Nesbitt theorem + +Q4958230: + title: Brauer–Siegel theorem + +Q4958233: + title: Brauer–Suzuki–Wall theorem + +Q4971339: + title: British flag theorem + +Q4975963: + title: Brown's representability theorem + +Q4979609: + title: Brun–Titchmarsh theorem + +Q4991415: + title: Chowla–Mordell theorem + +Q4998912: + title: Burke's theorem + +Q5001327: + title: Busemann's theorem + +Q5005143: + title: Bôcher's theorem + +Q5005965: + title: C-theorem + +Q5026435: + title: Cameron–Martin theorem + +Q5037752: + title: Carathéodory's existence theorem + +Q5041175: + title: Carleson–Jacobs theorem + +Q5042898: + title: Carlson's theorem + +Q5047043: + title: Cartan's theorems A and B + +Q5047046: + title: Brauer–Cartan–Hua theorem + +Q5047048: + title: Cartan–Hadamard theorem + +Q5047051: + title: Cartan–Kuranishi prolongation theorem + +Q5047052: + title: Cartan–Kähler theorem + +Q5049717: + title: Castelnuovo–de Franchis theorem + +Q5091194: + title: Cheng's eigenvalue comparison theorem + +Q5094298: + title: Chevalley's structure theorem + +Q5094305: + title: Chevalley–Shephard–Todd theorem + +Q5103861: + title: Choi's theorem on completely positive maps + +Q5125859: + title: Clapeyron's theorem + +Q5127710: + title: Clark–Ocone theorem + +Q5132839: + title: Clifford's circle theorems + +Q5132840: + title: Clifford's theorem on special divisors + +Q5136698: + title: Cluster decomposition theorem + +Q5139948: + title: Codd's theorem + +Q5141331: + title: Cohen structure theorem + +Q5155090: + title: Commutation theorem + +Q5157054: + title: Compression theorem + +Q5161245: + title: Conley–Zehnder theorem + +Q5163116: + title: Conservativity theorem + +Q5165492: + title: Continuous mapping theorem + +Q5166389: + title: Van Vleck's theorem + +Q5171652: + title: Corners theorem + +Q5172143: + title: Corona theorem + +Q5178114: + title: Courcelle's theorem + +Q5180651: + title: Craig's theorem + +Q5187911: + title: Crooks fluctuation theorem + +Q5188510: + title: Crossbar theorem + +Q5195996: + title: Curtis–Hedlund–Lyndon theorem + +Q5221089: + title: Danskin's theorem + +Q5242704: + title: Dawson–Gärtner theorem + +Q5244285: + title: de Bruijn's theorem + +Q5244361: + title: De Franchis theorem + +Q5251122: + title: Time hierarchy theorem + +Q5252320: + title: Lickorish twist theorem + +Q5278348: + title: Galvin's theorem + +Q5282059: + title: Harish–Chandra theorem + +Q5282247: + title: Disintegration theorem + +Q5295351: + title: Donaldson's theorem + +Q5296972: + title: Doob–Meyer decomposition theorem + +Q5299605: + title: Glivenko's theorem + +Q5311783: + title: Dudley's theorem + +Q5315060: + title: Dunford–Schwartz theorem + +Q5319505: + title: Zeilberger–Bressoud theorem + +Q5337931: + title: Edgeworth's limit theorem + +Q5362000: + title: Elitzur's theorem + +Q5384150: + title: Equal incircles theorem + +Q5385323: + title: Erdős–Gallai theorem + +Q5385324: + title: Erdős–Nagy theorem + +Q5385325: + title: Erdős–Pósa theorem + +Q5385326: + title: Erdős–Stone theorem + +Q5421989: + title: Exterior angle theorem + +Q5436274: + title: Farrell–Markushevich theorem + +Q5437947: + title: Fatou–Lebesgue theorem + +Q5438320: + title: Faustman–Ohlin theorem + +Q5443003: + title: Fenchel's duality theorem + +Q5443004: + title: Fenchel's theorem + +Q5443005: + title: Fenchel–Moreau theorem + +Q5445112: + title: Fernique's theorem + +Q5445311: + title: Ferrero–Washington theorem + +Q5447238: + title: Fieller's theorem + +Q5454965: + title: Fisher–Tippett–Gnedenko theorem + +Q5455495: + title: Fitting's theorem + +Q5456213: + title: Five circles theorem + +Q5463860: + title: Focal subgroup theorem + +Q5473576: + title: Foster's theorem + +Q5494134: + title: Fraňková–Helly selection theorem + +Q5498822: + title: Birkhoff's theorem + +Q5499906: + title: Shirshov–Witt theorem + +Q5501344: + title: Freidlin–Wentzell theorem + +Q5503689: + title: Bombieri–Friedlander–Iwaniec theorem + +Q5504427: + title: Friendship theorem + +Q5505098: + title: Frobenius determinant theorem + +Q5505114: + title: Froda's theorem + +Q5506929: + title: Fuchs's theorem + +Q5507285: + title: Fuglede's theorem + +Q5508180: + title: Full employment theorem + +Q5508482: + title: Fulton–Hansen connectedness theorem + +Q5508973: + title: Fundamental theorem of arbitrage-free pricing + +Q5530454: + title: Gell-Mann and Low theorem + +Q5552370: + title: Geroch's splitting theorem + +Q5566491: + title: Glaisher's theorem + +Q5567394: + title: Gleason's theorem + +Q5576268: + title: Goddard–Thorn theorem + +Q5579030: + title: Goldberg–Sachs theorem + +Q5580171: + title: Goldie's theorem + +Q5586067: + title: Gordon–Newell theorem + +Q5588002: + title: Gottesman–Knill theorem + +Q5597098: + title: Graph structure theorem + +Q5609384: + title: Grinberg's theorem + +Q5610188: + title: Gromov's compactness theorem + +Q5610190: + title: Gromov's compactness theorem + +Q5610718: + title: Grothendieck's connectedness theorem + +Q5612131: + title: Grunsky's theorem + +Q5612159: + title: Grunwald–Wang theorem + +Q5612871: + title: Grötzsch's theorem + +Q5638112: + title: Hadwiger's theorem + +Q5638886: + title: Hahn embedding theorem + +Q5643485: + title: Halpern–Läuchli theorem + +Q5645731: + title: Hammersley–Clifford theorem + +Q5656673: + title: Hardy–Littlewood tauberian theorem + +Q5656674: + title: Hardy–Ramanujan theorem + +Q5657833: + title: Harish–Chandra's regularity theorem + +Q5659675: + title: Harnack's curve theorem + +Q5675112: + title: Hartogs–Rosenthal theorem + +Q5680041: + title: Hasse–Arf theorem + +Q5697916B: + title: Waldhausen's theorem + +Q5697916A: + title: Reidemeister–Singer Theorem + +Q5697927: + title: Gross–Zagier theorem + +Q5709171: + title: Helly–Bray theorem + +Q5709377: + title: Helmholtz theorem (classical mechanics) + +Q5761142: + title: Hilbert's irreducibility theorem + +Q5874979: + title: Hobby–Rice theorem + +Q5876058: + title: Hodge index theorem + +Q5988409: + title: Identity theorem for Riemann surfaces + +Q6015420: + title: Increment theorem + +Q6059532: + title: Aronszajn–Smith theorem + +Q6086104: + title: Isomorphism extension theorem + +Q6119825: + title: Jacobson–Bourbaki theorem + +Q6276298: + title: Jordan's theorem (multiply transitive groups) + +Q6344729: + title: Kachurovskii's theorem + +Q6360586: + title: Kanamori–McAloon theorem + +Q6373327: + title: Karp–Lipton theorem + +Q6378596: + title: Katz–Lang finiteness theorem + +Q6379606: + title: Kawamata–Viehweg vanishing theorem + +Q6379715: + title: Kawasaki's theorem + +Q6387087: + title: Kempf–Ness theorem + +Q6400676: + title: Kharitonov's theorem + +Q6402928: + title: Lagrange reversion theorem + +Q6403282: + title: Lagrange's theorem + +Q6407842: + title: Killing–Hopf theorem + +Q6414200: + title: Kinoshita–Lee–Nauenberg theorem + +Q6421981: + title: Kneser's theorem + +Q6425088: + title: Kodaira embedding theorem + +Q6436753: + title: Krener's theorem + +Q6442276: + title: Kuiper's theorem + +Q6446396: + title: Kurosh subgroup theorem + +Q6455211: + title: Kōmura's theorem + +Q6456122: + title: L-balance theorem + +Q6471668: + title: Lafforgue's theorem + +Q6481192: + title: Lambek–Moser theorem + +Q6484331: + title: Landau prime ideal theorem + +Q6501470: + title: Lauricella's theorem + +Q6513990: + title: Lee Hwa Chung theorem + +Q6516611: + title: Lee–Yang theorem + +Q6516731: + title: Lefschetz hyperplane theorem + +Q6516736: + title: Lefschetz theorem on (1,1)-classes + +Q6528747: + title: Leray's theorem + +Q6528751: + title: Leray–Hirsch theorem + +Q6528849: + title: Lerner symmetry theorem + +Q6535568: + title: Levi's theorem + +Q6535741: + title: Levitzky's theorem + +Q6544508: + title: Lie–Palais theorem + +Q6701653: + title: Lukacs's proportion-sum independence theorem + +Q6707093: + title: Lyapunov–Malkin theorem + +Q6711207: + title: Lévy's modulus of continuity theorem + +Q6722024: + title: MacMahon Master theorem + +Q6733278: + title: Maharam's theorem + +Q6734194: + title: Mahler's compactness theorem + +Q6735549: + title: Maier's theorem + +Q6743217: + title: Malgrange preparation theorem + +Q6749789: + title: Manin–Drinfeld theorem + +Q6757284: + title: Marcinkiewicz theorem + +Q6760432: + title: Marginal value theorem + +Q6771536: + title: Markus−Yamabe theorem + +Q6777133: + title: Martingale representation theorem + +Q6783821: + title: Mason–Stothers theorem + +Q6795637: + title: Maximal ergodic theorem + +Q6795830: + title: Separating axis theorem + +Q6796056: + title: Maxwell's theorem + +Q6799039: + title: Mazur's control theorem + +Q6813106: + title: Mellin inversion theorem + +Q6859627: + title: Milliken's tree theorem + +Q6859648: + title: Milliken–Taylor theorem + +Q6860180: + title: Milman–Pettis theorem + +Q6867853: + title: Minkowski's second theorem + +Q6867867: + title: Minkowski–Hlawka theorem + +Q6867878: + title: Minlos's theorem + +Q6911202: + title: Moreau's theorem + +Q6914794: + title: Morton's theorem + +Q6925496: + title: Mountain pass theorem + +Q6927135: + title: Moving equilibrium theorem + +Q6935007: + title: Multiplicity-one theorem + +Q6935403: + title: Mumford vanishing theorem + +Q6957141: + title: Nachbin's theorem + +Q6967039: + title: Nash–Moser theorem + +Q7020025: + title: Newton's theorem about ovals + +Q7031618: + title: Nielsen realization problem + +Q7031621: + title: Nielsen fixed-point theorem + +Q7042768: + title: No-broadcasting theorem + +Q7042801: + title: No-trade theorem + +Q7043732: + title: Kuhn's theorem + +Q7045226: + title: No free lunch theorem + +Q7047379: + title: Noether's theorem on rationality for surfaces + +Q7098850: + title: Optical equivalence theorem + +Q7100033A: + title: Orbit theorem (Nagano–Sussmann) + +Q7100033B: + title: Rashevsky–Chow theorem + +Q7103669: + title: Ornstein theorem + +Q7106480: + title: Oseledec theorem + +Q7127466: + title: Paley's theorem + +Q7130794: + title: Pandya theorem + +Q7137494: + title: Paris–Harrington theorem + +Q7139570: + title: Parovicenko's theorem + +Q7140218: + title: Parthasarathy's theorem + +Q7160302: + title: Peeling theorem + +Q7160489: + title: Peetre theorem + +Q7160983: + title: Peixoto's theorem + +Q7190748: + title: Pickands–Balkema–de Haan theorem + +Q7200963: + title: Planar separator theorem + +Q7208500: + title: Poisson limit theorem + +Q7239984: + title: Preimage theorem + +Q7245073: + title: Principal axis theorem + +Q7249519: + title: Prokhorov's theorem + +Q7255475: + title: Pseudorandom generator theorem + +Q7269076: + title: No-deleting theorem + +Q7269106: + title: Quantum threshold theorem + +Q7269437: + title: Denjoy–Carleman theorem + +Q7269559: + title: Sylvester pentahedral theorem + +Q7272004: + title: Quillen–Suslin theorem + +Q7272898: + title: Quotient of subspace theorem + +Q7283856: + title: Raikov's theorem + +Q7288988: + title: Ramanujam vanishing theorem + +Q7295875: + title: Ratner's theorems + +Q7296106: + title: Rauch comparison theorem + +Q7307252: + title: Reflection theorem + +Q7308146: + title: Regev's theorem + +Q7309601: + title: Whitney–Graustein Theorem + +Q7310041: + title: Reider's theorem + +Q7312009: + title: Remmert–Stein theorem + +Q7318284: + title: Reversed compound agent theorem + +Q7322366: + title: Ribet's theorem + +Q7323144: + title: Rice–Shapiro theorem + +Q7333126: + title: Riemann–Roch theorem for surfaces + +Q7341043: + title: Robbins theorem + +Q7352955: + title: Robinson's joint consistency theorem + +Q7359997: + title: Rokhlin's theorem + +Q7396621: + title: Saccheri–Legendre theorem + +Q7430892: + title: Schaefer's dichotomy theorem + +Q7431298: + title: Schilder's theorem + +Q7431925: + title: Schnyder's theorem + +Q7432872: + title: Schreier refinement theorem + +Q7432915: + title: Schroeder–Bernstein theorem for measurable spaces + +Q7432918: + title: Schröder–Bernstein theorems for operator algebras + +Q7433034: + title: Great orthogonality theorem + +Q7433182: + title: Schwartz–Zippel theorem + +Q7433295: + title: Osterwalder–Schrader theorem + +Q7437564: + title: Scott core theorem + +Q7455422: + title: Swan's theorem + +Q7496252: + title: Shift theorem + +Q7510574: + title: Siegel–Walfisz theorem + +Q7512855: + title: Hirzebruch signature theorem + +Q7516736: + title: Silverman–Toeplitz theorem + +Q7524563: + title: Sinkhorn's theorem + +Q7525341: + title: Sion's minimax theorem + +Q7525845: + title: Sipser–Lautemann theorem + +Q7532192: + title: Siu's semicontinuity theorem + +Q7536097: + title: Skoda–El Mir theorem + +Q7536350: + title: Skorokhod's embedding theorem + +Q7572588: + title: Space hierarchy theorem + +Q7574438: + title: Specht's theorem + +Q7597316: + title: Stallings theorem about ends of groups + +Q7597317: + title: Stallings–Zeeman theorem + +Q7599389: + title: Stanley's reciprocity theorem + +Q7601243: + title: Star of David theorem + +Q7606897: + title: Steinitz theorem + +Q7606940: + title: Stein–Strömberg theorem + +Q7617407: + title: Stinespring factorization theorem + +Q7619060: + title: The duality theorem + +Q7619449: + title: Stone–von Neumann theorem + +Q7621715: + title: Strassmann's theorem + +Q7625164: + title: Structure theorem for Gaussian measures + +Q7632041: + title: Subspace theorem + +Q7644264: + title: Supersymmetry nonrenormalization theorems + +Q7660749: + title: Sylvester's determinant theorem + +Q7661308: + title: Symmetric hypergraph theorem + +Q7664013: + title: Sz.-Nagy's dilation theorem + +Q7686760: + title: Bang's theorem + +Q7782345: + title: Bertini's theorem + +Q7782346: + title: Theorem of the cube + +Q7782348: + title: Theorem of three moments + +Q7795825: + title: Thompson transitivity theorem + +Q7795828: + title: Thompson uniqueness theorem + +Q7809920: + title: Titchmarsh convolution theorem + +Q7818977: + title: Tomita's theorem + +Q7820874: + title: Tonelli's theorem + +Q7824894: + title: Topkis's theorem + +Q7825061: + title: Toponogov's theorem + +Q7825663: + title: Torelli theorem + +Q7827204: + title: Mazur's torsion theorem + +Q7841060: + title: Trichotomy theorem + +Q7845353: + title: Trombi–Varadarajan theorem + +Q7847268: + title: Trudinger's theorem + +Q7849521: + title: Tsen's theorem + +Q7853325: + title: Tunnell's theorem + +Q7856599: + title: Turán–Kubilius theorem + +Q7857350: + title: Tverberg's theorem + +Q7888360: + title: Structure theorem for finitely generated modules over a principal ideal domain + +Q7894110: + title: Universal approximation theorem + +Q7911648: + title: Valiant–Vazirani theorem + +Q7928688: + title: Vietoris–Begle mapping theorem + +Q7936914: + title: Vitali–Hahn–Saks theorem + +Q7941491: + title: Von Neumann's theorem + +Q7959587: + title: Wagner's theorem + +Q7966545: + title: Walter theorem + +Q7978735: + title: Weber's theorem + +Q7980241: + title: Weinberg–Witten theorem + +Q7996766: + title: Whitney extension theorem + +Q7996769: + title: Whitney immersion theorem + +Q7999144: + title: Wiener–Ikehara theorem + +Q7999797: + title: Beer's theorem + +Q8002487: + title: Wilkie's theorem + +Q8028529: + title: Witt's theorem + +Q8062800: + title: Z* theorem + +Q8063120: + title: ZJ theorem + +Q8064722: + title: Zahorski theorem + +Q8066611: + title: Kövari–Sós–Turán theorem + +Q8066795: + title: Zariski's main theorem + +Q8074796: + title: Zsigmondy's theorem + +Q8081891: + title: Śleszyński–Pringsheim theorem + +Q9993851: + title: Lax–Milgram theorem + +Q10369454: + title: Noether's second theorem + +Q10859514: + title: Vafa–Witten theorem + +Q10942247: + title: Hironaka theorem + +Q11352023: + title: Vitali convergence theorem + +Q11573495: + title: Plancherel theorem for spherical functions + +Q11704319: + title: Chern–Gauss–Bonnet theorem + +Q11722674: + title: Takens's theorem + +Q11883897: + title: Nagata's compactification theorem + +Q15080987: + title: Lie's third theorem + +Q15813392: + title: Barban–Davenport–Halberstam theorem + +Q15830473: + title: Morley's categoricity theorem + +Q15844093: + title: Van Schooten's theorem + +Q15859323: + title: Anne's theorem + +Q15872491: + title: Newton's theorem (quadrilateral) + +Q15895894: + title: Finsler–Hadwiger theorem + +Q16251580: + title: Matiyasevich's theorem + +Q16680059: + title: Strong perfect graph theorem + +Q16978447: + title: Grauert–Riemenschneider vanishing theorem + +Q17001601: + title: 2-factor theorem + +Q17002391: + title: Arrow-Lind theorem + +Q17003552: + title: Alexandrov's uniqueness theorem + +Q17005116: + title: Birkhoff's representation theorem + +Q17008559: + title: Davenport–Schmidt theorem + +Q17017973: + title: Godunov's theorem + +Q17018210: + title: Golod–Shafarevich theorem + +Q17019684: + title: Grushko theorem + +Q17029787: + title: Higman's embedding theorem + +Q17080564: + title: Zariski's connectedness theorem + +Q17081508: + title: Bing metrization theorem + +Q17082552: + title: Mutual fund separation theorem + +Q17089994X: + title: Tikhonov fixed-point theorem + +Q17089994: + title: Fixed-point theorems in infinite-dimensional spaces + +Q17098075: + title: Jurkat–Richert theorem + +Q17098298: + title: Jacobson density theorem + +Q17098379: + title: Kaplansky's theorem on quadratic forms + +Q17101806: + title: Intersection theorem + +Q17102744: + title: Riemann–Roch theorem for smooth manifolds + +Q17103352: + title: Maximum power theorem + +Q17104025: + title: Riemann singularity theorem + +Q17104863: + title: Nielsen–Ninomiya theorem + +Q17125544: + title: Lickorish–Wallace theorem + +Q18205730: + title: Byers–Yang theorem + +Q18206032: + title: Cartan's theorem + +Q18206266: + title: Euclid–Euler theorem + +Q18630480: + title: Euler's quadrilateral theorem + +Q19323571: + title: Chomsky–Schützenberger enumeration theorem + +Q19779530: + title: Thomsen's theorem + +Q20278711: + title: Cramer's theorem (algebraic curves) + +Q20971632: + title: Lie's theorem + +Q22952648: + title: Uncountability of the continuum + +Q25099402: + title: Kolmogorov–Arnold representation theorem + +Q25303622: + title: Browder–Minty theorem + +Q25304227: + title: Arakelyan's theorem + +Q25304301: + title: Bregman–Minc inequality + +Q25345219: + title: BBD decomposition theorem + +Q25378690: + title: Ionescu-Tulcea theorem + +Q26877569: + title: Furry's theorem + +Q28194853: + title: Lovelock's theorem + +Q28458131: + title: Glivenko's theorem + +Q31838822: + title: Kirchberger's theorem + +Q48996535: + title: Solèr's theorem + +Q48998319: + title: Frobenius reciprocity theorem + +Q55647729: + title: Cramér's theorem (large deviations) + +Q56291669: + title: Alspach's theorem + +Q60681777: + title: Constant chord theorem + +Q61163749: + title: Theorem of the gnomon + +Q62051012: + title: Behrend's theorem + +Q96375449: + title: Conway circle theorem + +Q96377355: + title: Erdős–Dushnik–Miller theorem + +Q97359729: + title: Stahl's theorem + +Q104841721: + title: Jacobson–Morozov theorem + +Q104871594: + title: Joubert's theorem + +Q105222412: + title: Five color theorem + +Q107535899: + title: Bochner's tube theorem + +Q107547910: + title: Malgrange–Zerner theorem + +Q107709489: + title: Büchi-Elgot-Trakhtenbrot theorem + +Q107710112: + title: Cantor's isomorphism theorem + +Q110921617: + title: Feferman–Vaught theorem + +Q111982312: + title: Friedberg–Muchnik theorem + +Q112659261: + title: Gamas's Theorem + +Q112740521: + title: Netto's theorem diff --git a/scripts/check-yaml.lean b/scripts/check-yaml.lean index 5d785ad1a5928..37da03fd90436 100644 --- a/scripts/check-yaml.lean +++ b/scripts/check-yaml.lean @@ -7,7 +7,7 @@ import Lean.Util.SearchPath import Mathlib.Lean.CoreM import Mathlib.Tactic.ToExpr -/-! # Script to check `undergrad.yaml`, `overview.yaml`, and `100.yaml` +/-! # Script to check `undergrad.yaml`, `overview.yaml`, `100.yaml` and `1000.yaml` This assumes `yaml_check.py` has first translated these to `json` files. @@ -23,7 +23,7 @@ def readJsonFile (α) [FromJson α] (path : System.FilePath) : IO α := do let _ : MonadExceptOf String IO := ⟨throw ∘ IO.userError, fun x _ => x⟩ liftExcept <| fromJson? <|← liftExcept <| Json.parse <|← IO.FS.readFile path -def databases : List String := ["undergrad", "overview", "100"] +def databases : List String := ["undergrad", "overview", "100", "1000"] def processDb (decls : ConstMap) : String → IO Bool | file => do diff --git a/scripts/yaml_check.py b/scripts/yaml_check.py index 4e05448995288..7605122447a6d 100644 --- a/scripts/yaml_check.py +++ b/scripts/yaml_check.py @@ -1,9 +1,9 @@ """ -This file reads in the three yaml files, and translates them to simpler json files +This file reads in the four yaml files, and translates them to simpler json files that are easier to process in Lean. Usage: - python3 yaml_check.py <100.yaml> + python3 yaml_check.py <100.yaml> <1000.yaml> (Each is the path to a yaml file containing information about the respective class of theorems. The order of these files is important.) @@ -38,17 +38,20 @@ def print_list(fn: str, pairs: List[Tuple[str, str]]) -> None: out.write(f'{id}\n{val.strip()}\n\n') hundred_yaml = sys.argv[1] -overview_yaml = sys.argv[2] -undergrad_yaml = sys.argv[3] +thousand_yaml = sys.argv[2] +overview_yaml = sys.argv[3] +undergrad_yaml = sys.argv[4] with open(hundred_yaml, 'r', encoding='utf8') as hy: hundred = yaml.safe_load(hy) +with open(thousand_yaml, 'r', encoding='utf8') as hy: + thousand = yaml.safe_load(hy) with open(overview_yaml, 'r', encoding='utf8') as hy: overview = yaml.safe_load(hy) with open(undergrad_yaml, 'r', encoding='utf8') as hy: undergrad = yaml.safe_load(hy) -hundred_decls:List[Tuple[str, str]] = [] +hundred_decls: List[Tuple[str, str]] = [] for index, entry in hundred.items(): title = entry['title'] @@ -59,6 +62,16 @@ def print_list(fn: str, pairs: List[Tuple[str, str]]) -> None: raise ValueError(f"For key {index} ({title}): did you mean `decl` instead of `decls`?") hundred_decls = hundred_decls + [(f'{index} {title}', d) for d in entry['decls']] +thousand_decls: List[Tuple[str, str]] = [] +for index, entry in thousand.items(): + title = entry['title'] + if 'decl' in entry: + thousand_decls.append((f'{index} {title}', entry['decl'])) + elif 'decls' in entry: + if not isinstance(entry['decls'], list): + raise ValueError(f"For key {index} ({title}): did you mean `decl` instead of `decls`?") + thousand_decls = thousand_decls + [(f'{index} {title}', d) for d in entry['decls']] + overview_decls = tiered_extract(overview) assert all(len(n) == 3 for n, _ in overview_decls) overview_decls = flatten_names(overview_decls) @@ -69,6 +82,8 @@ def print_list(fn: str, pairs: List[Tuple[str, str]]) -> None: with open('100.json', 'w', encoding='utf8') as f: json.dump(hundred_decls, f) +with open('1000.json', 'w', encoding='utf8') as f: + json.dump(thousand_decls, f) with open('overview.json', 'w', encoding='utf8') as f: json.dump(overview_decls, f) with open('undergrad.json', 'w', encoding='utf8') as f: From a2702bb7b4b5c1a56c5fe023f75b122dfd89fe66 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Mon, 9 Dec 2024 09:39:54 +0000 Subject: [PATCH 587/829] feat(TensorProduct): change of base ring in tensor product of algebras (#19670) A follow up to #18744, adding more contents about tensor product of algebras. In `LinearAlgebra/TensorProduct/Basic.lean`: change the argument order of `mapOfCompatibleSMul` from `R M N A` to the more natural `R A M N`; add `lidOfCompatibleSMul`. In `RingTheory/TensorProduct/Basic.lean`: add a CompatibleSMul section for algebras completely analogous to the CompatibleSMul section in `LinearAlgebra/TensorProduct/Basic.lean` added in #18744. In `RingTheory/Localization/BaseChange.lean`: promote various examples to defs, and generalize `tensorSelfEquiv` to `algebraLid`. --- .../LinearAlgebra/TensorProduct/Basic.lean | 99 ++++++++++--------- .../RingTheory/Localization/BaseChange.lean | 37 ++++--- Mathlib/RingTheory/TensorProduct/Basic.lean | 55 +++++++++++ 3 files changed, 135 insertions(+), 56 deletions(-) diff --git a/Mathlib/LinearAlgebra/TensorProduct/Basic.lean b/Mathlib/LinearAlgebra/TensorProduct/Basic.lean index 0a966b059726d..75d5ed0bc25c0 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Basic.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Basic.lean @@ -45,7 +45,7 @@ section Semiring variable {R : Type*} [CommSemiring R] variable {R' : Type*} [Monoid R'] variable {R'' : Type*} [Semiring R''] -variable {M : Type*} {N : Type*} {P : Type*} {Q : Type*} {S : Type*} {T : Type*} +variable {A M N P Q S T : Type*} variable [AddCommMonoid M] [AddCommMonoid N] [AddCommMonoid P] variable [AddCommMonoid Q] [AddCommMonoid S] [AddCommMonoid T] variable [Module R M] [Module R N] [Module R Q] [Module R S] [Module R T] @@ -552,49 +552,6 @@ theorem lift_compr₂ (g : P →ₗ[R] Q) : lift (f.compr₂ g) = g.comp (lift f theorem lift_mk_compr₂ (f : M ⊗ N →ₗ[R] P) : lift ((mk R M N).compr₂ f) = f := by rw [lift_compr₂ f, lift_mk, LinearMap.comp_id] -section ScalarTower - -variable (R M N A) [CommSemiring A] [Module A M] [Module A N] [SMulCommClass R A M] - [CompatibleSMul R A M N] - -/-- If M and N are both R- and A-modules and their actions on them commute, -and if the A-action on `M ⊗[R] N` can switch between the two factors, then there is a -canonical A-linear map from `M ⊗[A] N` to `M ⊗[R] N`. -/ -def mapOfCompatibleSMul : M ⊗[A] N →ₗ[A] M ⊗[R] N := - lift - { toFun := fun m ↦ - { __ := mk R M N m - map_smul' := fun _ _ ↦ (smul_tmul _ _ _).symm } - map_add' := fun _ _ ↦ LinearMap.ext <| by simp - map_smul' := fun _ _ ↦ rfl } - -@[simp] theorem mapOfCompatibleSMul_tmul (m n) : mapOfCompatibleSMul R M N A (m ⊗ₜ n) = m ⊗ₜ n := - rfl - -attribute [local instance] SMulCommClass.symm - -/-- `mapOfCompatibleSMul R M N A` is also A-linear. -/ -def mapOfCompatibleSMul' : M ⊗[A] N →ₗ[R] M ⊗[R] N where - __ := mapOfCompatibleSMul R M N A - map_smul' _ x := x.induction_on (map_zero _) (fun _ _ ↦ by simp [smul_tmul']) - fun _ _ h h' ↦ by simpa using congr($h + $h') - -theorem mapOfCompatibleSMul_surjective : Function.Surjective (mapOfCompatibleSMul R M N A) := - fun x ↦ x.induction_on (⟨0, map_zero _⟩) (fun m n ↦ ⟨_, mapOfCompatibleSMul_tmul ..⟩) - fun _ _ ⟨x, hx⟩ ⟨y, hy⟩ ↦ ⟨x + y, by simpa using congr($hx + $hy)⟩ - -/-- If the R- and A-actions on M and N satisfy `CompatibleSMul` both ways, -then `M ⊗[A] N` is canonically isomorphic to `M ⊗[R] N`. -/ -def equivOfCompatibleSMul [CompatibleSMul A R M N] : M ⊗[A] N ≃ₗ[A] M ⊗[R] N where - __ := mapOfCompatibleSMul R M N A - invFun := mapOfCompatibleSMul A M N R - left_inv x := x.induction_on (map_zero _) (fun _ _ ↦ rfl) - fun _ _ h h' ↦ by simpa using congr($h + $h') - right_inv x := x.induction_on (map_zero _) (fun _ _ ↦ rfl) - fun _ _ h h' ↦ by simpa using congr($h + $h') - -end ScalarTower - /-- This used to be an `@[ext]` lemma, but it fails very slowly when the `ext` tactic tries to apply it in some cases, notably when one wants to show equality of two linear maps. The `@[ext]` attribute is now added locally where it is needed. Using this as the `@[ext]` lemma instead of @@ -756,6 +713,60 @@ variable (R) in theorem lid_eq_rid : TensorProduct.lid R R = TensorProduct.rid R R := LinearEquiv.toLinearMap_injective <| ext' mul_comm +section CompatibleSMul + +variable (R A M N) [CommSemiring A] [Module A M] [Module A N] [SMulCommClass R A M] + [CompatibleSMul R A M N] + +/-- If M and N are both R- and A-modules and their actions on them commute, +and if the A-action on `M ⊗[R] N` can switch between the two factors, then there is a +canonical A-linear map from `M ⊗[A] N` to `M ⊗[R] N`. -/ +def mapOfCompatibleSMul : M ⊗[A] N →ₗ[A] M ⊗[R] N := + lift + { toFun := fun m ↦ + { __ := mk R M N m + map_smul' := fun _ _ ↦ (smul_tmul _ _ _).symm } + map_add' := fun _ _ ↦ LinearMap.ext <| by simp + map_smul' := fun _ _ ↦ rfl } + +@[simp] theorem mapOfCompatibleSMul_tmul (m n) : mapOfCompatibleSMul R A M N (m ⊗ₜ n) = m ⊗ₜ n := + rfl + +theorem mapOfCompatibleSMul_surjective : Function.Surjective (mapOfCompatibleSMul R A M N) := + fun x ↦ x.induction_on (⟨0, map_zero _⟩) (fun m n ↦ ⟨_, mapOfCompatibleSMul_tmul ..⟩) + fun _ _ ⟨x, hx⟩ ⟨y, hy⟩ ↦ ⟨x + y, by simpa using congr($hx + $hy)⟩ + +attribute [local instance] SMulCommClass.symm + +/-- `mapOfCompatibleSMul R A M N` is also R-linear. -/ +def mapOfCompatibleSMul' : M ⊗[A] N →ₗ[R] M ⊗[R] N where + __ := mapOfCompatibleSMul R A M N + map_smul' _ x := x.induction_on (map_zero _) (fun _ _ ↦ by simp [smul_tmul']) + fun _ _ h h' ↦ by simpa using congr($h + $h') + +/-- If the R- and A-actions on M and N satisfy `CompatibleSMul` both ways, +then `M ⊗[A] N` is canonically isomorphic to `M ⊗[R] N`. -/ +def equivOfCompatibleSMul [CompatibleSMul A R M N] : M ⊗[A] N ≃ₗ[A] M ⊗[R] N where + __ := mapOfCompatibleSMul R A M N + invFun := mapOfCompatibleSMul A R M N + left_inv x := x.induction_on (map_zero _) (fun _ _ ↦ rfl) + fun _ _ h h' ↦ by simpa using congr($h + $h') + right_inv x := x.induction_on (map_zero _) (fun _ _ ↦ rfl) + fun _ _ h h' ↦ by simpa using congr($h + $h') + +omit [SMulCommClass R A M] + +variable [Module R A] [SMulCommClass R A A] [CompatibleSMul R A A M] [CompatibleSMul A R A M] + +/-- If the R- and A- action on A and M satisfy `CompatibleSMul` both ways, +then `A ⊗[R] M` is canonically isomorphic to `M`. -/ +def lidOfCompatibleSMul : A ⊗[R] M ≃ₗ[A] M := + (equivOfCompatibleSMul R A A M).symm ≪≫ₗ TensorProduct.lid _ _ + +theorem lidOfCompatibleSMul_tmul (a m) : lidOfCompatibleSMul R A M (a ⊗ₜ[R] m) = a • m := rfl + +end CompatibleSMul + open LinearMap section diff --git a/Mathlib/RingTheory/Localization/BaseChange.lean b/Mathlib/RingTheory/Localization/BaseChange.lean index 62ba26f974341..bbddee1b93312 100644 --- a/Mathlib/RingTheory/Localization/BaseChange.lean +++ b/Mathlib/RingTheory/Localization/BaseChange.lean @@ -54,33 +54,46 @@ namespace IsLocalization include S open TensorProduct Algebra.TensorProduct -variable (M₁ M₂) [AddCommMonoid M₁] [AddCommMonoid M₂] [Module R M₁] [Module R M₂] +variable (M₁ M₂ B C) [AddCommMonoid M₁] [AddCommMonoid M₂] [Module R M₁] [Module R M₂] [Module A M₁] [Module A M₂] [IsScalarTower R A M₁] [IsScalarTower R A M₂] + [Semiring B] [Algebra R B] [Algebra A B] [IsScalarTower R A B] + [Semiring C] [Algebra R C] [Algebra A C] [IsScalarTower R A C] theorem tensorProduct_compatibleSMul : CompatibleSMul R A M₁ M₂ where smul_tmul a _ _ := by - obtain ⟨r, s, rfl⟩ := IsLocalization.mk'_surjective S a + obtain ⟨r, s, rfl⟩ := mk'_surjective S a rw [← (map_units A s).smul_left_cancel] simp_rw [algebraMap_smul, smul_tmul', ← smul_assoc, smul_tmul, ← smul_assoc, smul_mk'_self, algebraMap_smul, smul_tmul] -noncomputable example : M₁ ⊗[A] M₂ ≃ₗ[A] M₁ ⊗[R] M₂ := +/-- If `A` is a localization of `R`, tensoring two `A`-modules over `A` is the same as +tensoring them over `R`. -/ +noncomputable def moduleTensorEquiv : M₁ ⊗[A] M₂ ≃ₗ[A] M₁ ⊗[R] M₂ := have := tensorProduct_compatibleSMul S A M₁ M₂ - equivOfCompatibleSMul R M₁ M₂ A + equivOfCompatibleSMul R A M₁ M₂ -noncomputable example : A ⊗[R] M₁ ≃ₗ[A] M₁ := +/-- If `A` is a localization of `R`, tensoring an `A`-module with `A` over `R` does nothing. -/ +noncomputable def moduleLid : A ⊗[R] M₁ ≃ₗ[A] M₁ := have := tensorProduct_compatibleSMul S A A M₁ - (equivOfCompatibleSMul R A M₁ A).symm ≪≫ₗ TensorProduct.lid _ _ + (equivOfCompatibleSMul R A A M₁).symm ≪≫ₗ TensorProduct.lid _ _ -/-- If A is a localization of a commutative ring R, the tensor product of A with A over R is -canonically isomorphic as A-algebras to A itself. -/ -noncomputable def tensorSelfAlgEquiv : A ⊗[R] A ≃ₐ[A] A := - have := tensorProduct_compatibleSMul S A A A - lmulEquiv R A +/-- If `A` is a localization of `R`, tensoring two `A`-algebras over `A` is the same as +tensoring them over `R`. -/ +noncomputable def algebraTensorEquiv : B ⊗[A] C ≃ₐ[A] B ⊗[R] C := + have := tensorProduct_compatibleSMul S A B C + Algebra.TensorProduct.equivOfCompatibleSMul R A B C + +/-- If `A` is a localization of `R`, tensoring an `A`-algebra with `A` over `R` does nothing. -/ +noncomputable def algebraLid : A ⊗[R] B ≃ₐ[A] B := + have := tensorProduct_compatibleSMul S A A B + Algebra.TensorProduct.lidOfCompatibleSMul R A B + +@[deprecated (since := "2024-12-01")] alias tensorSelfAlgEquiv := algebraLid set_option linter.docPrime false in theorem bijective_linearMap_mul' : Function.Bijective (LinearMap.mul' R A) := - (tensorSelfAlgEquiv S A).bijective + have := tensorProduct_compatibleSMul S A A A + (Algebra.TensorProduct.lmulEquiv R A).bijective end IsLocalization diff --git a/Mathlib/RingTheory/TensorProduct/Basic.lean b/Mathlib/RingTheory/TensorProduct/Basic.lean index 857c189049c66..7eeec368f0ac8 100644 --- a/Mathlib/RingTheory/TensorProduct/Basic.lean +++ b/Mathlib/RingTheory/TensorProduct/Basic.lean @@ -807,6 +807,53 @@ variable {A} in @[simp] theorem rid_symm_apply (a : A) : (TensorProduct.rid R S A).symm a = a ⊗ₜ 1 := rfl +section CompatibleSMul + +variable (R S A B : Type*) [CommSemiring R] [CommSemiring S] [Semiring A] [Semiring B] +variable [Algebra R A] [Algebra R B] [Algebra S A] [Algebra S B] +variable [SMulCommClass R S A] [CompatibleSMul R S A B] + +/-- If A and B are both R- and S-algebras and their actions on them commute, +and if the S-action on `A ⊗[R] B` can switch between the two factors, then there is a +canonical S-algebra homomorphism from `A ⊗[S] B` to `A ⊗[R] B`. -/ +def mapOfCompatibleSMul : A ⊗[S] B →ₐ[S] A ⊗[R] B := + .ofLinearMap (_root_.TensorProduct.mapOfCompatibleSMul R S A B) rfl fun x ↦ + x.induction_on (by simp) (fun _ _ y ↦ y.induction_on (by simp) (by simp) + fun _ _ h h' ↦ by simp only [mul_add, map_add, h, h']) + fun _ _ h h' _ ↦ by simp only [add_mul, map_add, h, h'] + +@[simp] theorem mapOfCompatibleSMul_tmul (m n) : mapOfCompatibleSMul R S A B (m ⊗ₜ n) = m ⊗ₜ n := + rfl + +theorem mapOfCompatibleSMul_surjective : Function.Surjective (mapOfCompatibleSMul R S A B) := + _root_.TensorProduct.mapOfCompatibleSMul_surjective R S A B + +attribute [local instance] SMulCommClass.symm + +/-- `mapOfCompatibleSMul R S A B` is also A-linear. -/ +def mapOfCompatibleSMul' : A ⊗[S] B →ₐ[R] A ⊗[R] B := + .ofLinearMap (_root_.TensorProduct.mapOfCompatibleSMul' R S A B) rfl + (map_mul <| mapOfCompatibleSMul R S A B) + +/-- If the R- and S-actions on A and B satisfy `CompatibleSMul` both ways, +then `A ⊗[S] B` is canonically isomorphic to `A ⊗[R] B`. -/ +def equivOfCompatibleSMul [CompatibleSMul S R A B] : A ⊗[S] B ≃ₐ[S] A ⊗[R] B where + __ := mapOfCompatibleSMul R S A B + invFun := mapOfCompatibleSMul S R A B + __ := _root_.TensorProduct.equivOfCompatibleSMul R S A B + +variable [Algebra R S] [CompatibleSMul R S S A] [CompatibleSMul S R S A] +omit [SMulCommClass R S A] + +/-- If the R- and S- action on S and A satisfy `CompatibleSMul` both ways, +then `S ⊗[R] A` is canonically isomorphic to `A`. -/ +def lidOfCompatibleSMul : S ⊗[R] A ≃ₐ[S] A := + (equivOfCompatibleSMul R S S A).symm.trans (TensorProduct.lid _ _) + +theorem lidOfCompatibleSMul_tmul (s a) : lidOfCompatibleSMul R S A (s ⊗ₜ[R] a) = s • a := rfl + +end CompatibleSMul + section variable (B) @@ -1054,6 +1101,10 @@ def lmul'' : S ⊗[R] S →ₐ[S] S := fun x y hx hy ↦ by simp_all [hx, hy, mul_add] } (fun a₁ a₂ b₁ b₂ => by simp [mul_mul_mul_comm]) <| by simp +theorem lmul''_eq_lid_comp_mapOfCompatibleSMul : + lmul'' R = (TensorProduct.lid S S).toAlgHom.comp (mapOfCompatibleSMul' _ _ _ _) := by + ext; rfl + /-- `LinearMap.mul'` as an `AlgHom` over the base ring. -/ def lmul' : S ⊗[R] S →ₐ[R] S := (lmul'' R).restrictScalars R @@ -1083,6 +1134,10 @@ def lmulEquiv [CompatibleSMul R S S S] : S ⊗[R] S ≃ₐ[S] S := rw [mul_comm, ← smul_eq_mul, smul_tmul, smul_eq_mul, mul_one]) fun _ _ hx hy ↦ by simp_all [hx, hy, add_tmul] +theorem lmulEquiv_eq_lidOfCompatibleSMul [CompatibleSMul R S S S] : + lmulEquiv R S = lidOfCompatibleSMul R S S := + AlgEquiv.coe_algHom_injective <| by ext; rfl + /-- If `S` is commutative, for a pair of morphisms `f : A →ₐ[R] S`, `g : B →ₐ[R] S`, We obtain a map `A ⊗[R] B →ₐ[R] S` that commutes with `f`, `g` via `a ⊗ b ↦ f(a) * g(b)`. From b405b4d82c448946e73137f88260686e6ed97a3b Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 9 Dec 2024 09:39:56 +0000 Subject: [PATCH 588/829] chore(Analysis/BoundedVariation): split file (#19784) Fixes #4862 --- Mathlib.lean | 1 + Mathlib/Analysis/BoundedVariation.lean | 766 +---------------- Mathlib/Analysis/ConstantSpeed.lean | 5 +- .../EMetricSpace/BoundedVariation.lean | 777 ++++++++++++++++++ 4 files changed, 788 insertions(+), 761 deletions(-) create mode 100644 Mathlib/Topology/EMetricSpace/BoundedVariation.lean diff --git a/Mathlib.lean b/Mathlib.lean index 4ec2f1daeaf33..e287ba642ed7e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -5158,6 +5158,7 @@ import Mathlib.Topology.DerivedSet import Mathlib.Topology.DiscreteQuotient import Mathlib.Topology.DiscreteSubset import Mathlib.Topology.EMetricSpace.Basic +import Mathlib.Topology.EMetricSpace.BoundedVariation import Mathlib.Topology.EMetricSpace.Defs import Mathlib.Topology.EMetricSpace.Diam import Mathlib.Topology.EMetricSpace.Lipschitz diff --git a/Mathlib/Analysis/BoundedVariation.lean b/Mathlib/Analysis/BoundedVariation.lean index 4a2fafb5a8000..74167269753e8 100644 --- a/Mathlib/Analysis/BoundedVariation.lean +++ b/Mathlib/Analysis/BoundedVariation.lean @@ -7,784 +7,32 @@ import Mathlib.Analysis.Calculus.FDeriv.Add import Mathlib.Analysis.Calculus.FDeriv.Equiv import Mathlib.Analysis.Calculus.FDeriv.Prod import Mathlib.Analysis.Calculus.Monotone -import Mathlib.Data.Set.Function -import Mathlib.Algebra.Group.Basic -import Mathlib.Tactic.WLOG +import Mathlib.Topology.EMetricSpace.BoundedVariation /-! -# Functions of bounded variation +# Almost everywhere differentiability of functions with locally bounded variation -We study functions of bounded variation. In particular, we show that a bounded variation function -is a difference of monotone functions, and differentiable almost everywhere. This implies that -Lipschitz functions from the real line into finite-dimensional vector space are also differentiable -almost everywhere. +In this file we show that a bounded variation function is differentiable almost everywhere. +This implies that Lipschitz functions from the real line into finite-dimensional vector space +are also differentiable almost everywhere. ## Main definitions and results -* `eVariationOn f s` is the total variation of the function `f` on the set `s`, in `ℝ≥0∞`. -* `BoundedVariationOn f s` registers that the variation of `f` on `s` is finite. -* `LocallyBoundedVariationOn f s` registers that `f` has finite variation on any compact - subinterval of `s`. -* `variationOnFromTo f s a b` is the signed variation of `f` on `s ∩ Icc a b`, converted to `ℝ`. - -* `eVariationOn.Icc_add_Icc` states that the variation of `f` on `[a, c]` is the sum of its - variations on `[a, b]` and `[b, c]`. -* `LocallyBoundedVariationOn.exists_monotoneOn_sub_monotoneOn` proves that a function - with locally bounded variation is the difference of two monotone functions. -* `LipschitzWith.locallyBoundedVariationOn` shows that a Lipschitz function has locally - bounded variation. * `LocallyBoundedVariationOn.ae_differentiableWithinAt` shows that a bounded variation function into a finite dimensional real vector space is differentiable almost everywhere. * `LipschitzOnWith.ae_differentiableWithinAt` is the same result for Lipschitz functions. We also give several variations around these results. -## Implementation - -We define the variation as an extended nonnegative real, to allow for infinite variation. This makes -it possible to use the complete linear order structure of `ℝ≥0∞`. The proofs would be much -more tedious with an `ℝ`-valued or `ℝ≥0`-valued variation, since one would always need to check -that the sets one uses are nonempty and bounded above as these are only conditionally complete. -/ - -open scoped NNReal ENNReal Topology UniformConvergence - +open scoped NNReal ENNReal Topology open Set MeasureTheory Filter -- Porting note: sectioned variables because a `wlog` was broken due to extra variables in context variable {α : Type*} [LinearOrder α] {E : Type*} [PseudoEMetricSpace E] -/-- The (extended real valued) variation of a function `f` on a set `s` inside a linear order is -the supremum of the sum of `edist (f (u (i+1))) (f (u i))` over all finite increasing -sequences `u` in `s`. -/ -noncomputable def eVariationOn (f : α → E) (s : Set α) : ℝ≥0∞ := - ⨆ p : ℕ × { u : ℕ → α // Monotone u ∧ ∀ i, u i ∈ s }, - ∑ i ∈ Finset.range p.1, edist (f (p.2.1 (i + 1))) (f (p.2.1 i)) - -/-- A function has bounded variation on a set `s` if its total variation there is finite. -/ -def BoundedVariationOn (f : α → E) (s : Set α) := - eVariationOn f s ≠ ∞ - -/-- A function has locally bounded variation on a set `s` if, given any interval `[a, b]` with -endpoints in `s`, then the function has finite variation on `s ∩ [a, b]`. -/ -def LocallyBoundedVariationOn (f : α → E) (s : Set α) := - ∀ a b, a ∈ s → b ∈ s → BoundedVariationOn f (s ∩ Icc a b) - -/-! ## Basic computations of variation -/ - -namespace eVariationOn - -theorem nonempty_monotone_mem {s : Set α} (hs : s.Nonempty) : - Nonempty { u // Monotone u ∧ ∀ i : ℕ, u i ∈ s } := by - obtain ⟨x, hx⟩ := hs - exact ⟨⟨fun _ => x, fun i j _ => le_rfl, fun _ => hx⟩⟩ - -theorem eq_of_edist_zero_on {f f' : α → E} {s : Set α} (h : ∀ ⦃x⦄, x ∈ s → edist (f x) (f' x) = 0) : - eVariationOn f s = eVariationOn f' s := by - dsimp only [eVariationOn] - congr 1 with p : 1 - congr 1 with i : 1 - rw [edist_congr_right (h <| p.snd.prop.2 (i + 1)), edist_congr_left (h <| p.snd.prop.2 i)] - -theorem eq_of_eqOn {f f' : α → E} {s : Set α} (h : EqOn f f' s) : - eVariationOn f s = eVariationOn f' s := - eq_of_edist_zero_on fun x xs => by rw [h xs, edist_self] - -theorem sum_le (f : α → E) {s : Set α} (n : ℕ) {u : ℕ → α} (hu : Monotone u) (us : ∀ i, u i ∈ s) : - (∑ i ∈ Finset.range n, edist (f (u (i + 1))) (f (u i))) ≤ eVariationOn f s := - le_iSup_of_le ⟨n, u, hu, us⟩ le_rfl - -theorem sum_le_of_monotoneOn_Icc (f : α → E) {s : Set α} {m n : ℕ} {u : ℕ → α} - (hu : MonotoneOn u (Icc m n)) (us : ∀ i ∈ Icc m n, u i ∈ s) : - (∑ i ∈ Finset.Ico m n, edist (f (u (i + 1))) (f (u i))) ≤ eVariationOn f s := by - rcases le_total n m with hnm | hmn - · simp [Finset.Ico_eq_empty_of_le hnm] - let π := projIcc m n hmn - let v i := u (π i) - calc - ∑ i ∈ Finset.Ico m n, edist (f (u (i + 1))) (f (u i)) - = ∑ i ∈ Finset.Ico m n, edist (f (v (i + 1))) (f (v i)) := - Finset.sum_congr rfl fun i hi ↦ by - rw [Finset.mem_Ico] at hi - simp only [v, π, projIcc_of_mem hmn ⟨hi.1, hi.2.le⟩, - projIcc_of_mem hmn ⟨hi.1.trans i.le_succ, hi.2⟩] - _ ≤ ∑ i ∈ Finset.range n, edist (f (v (i + 1))) (f (v i)) := - Finset.sum_mono_set _ (Nat.Iio_eq_range ▸ Finset.Ico_subset_Iio_self) - _ ≤ eVariationOn f s := - sum_le _ _ (fun i j h ↦ hu (π i).2 (π j).2 (monotone_projIcc hmn h)) fun i ↦ us _ (π i).2 - -theorem sum_le_of_monotoneOn_Iic (f : α → E) {s : Set α} {n : ℕ} {u : ℕ → α} - (hu : MonotoneOn u (Iic n)) (us : ∀ i ≤ n, u i ∈ s) : - (∑ i ∈ Finset.range n, edist (f (u (i + 1))) (f (u i))) ≤ eVariationOn f s := by - simpa using sum_le_of_monotoneOn_Icc f (m := 0) (hu.mono Icc_subset_Iic_self) fun i hi ↦ us i hi.2 - -theorem mono (f : α → E) {s t : Set α} (hst : t ⊆ s) : eVariationOn f t ≤ eVariationOn f s := by - apply iSup_le _ - rintro ⟨n, ⟨u, hu, ut⟩⟩ - exact sum_le f n hu fun i => hst (ut i) - -theorem _root_.BoundedVariationOn.mono {f : α → E} {s : Set α} (h : BoundedVariationOn f s) - {t : Set α} (ht : t ⊆ s) : BoundedVariationOn f t := - ne_top_of_le_ne_top h (eVariationOn.mono f ht) - -theorem _root_.BoundedVariationOn.locallyBoundedVariationOn {f : α → E} {s : Set α} - (h : BoundedVariationOn f s) : LocallyBoundedVariationOn f s := fun _ _ _ _ => - h.mono inter_subset_left - -theorem edist_le (f : α → E) {s : Set α} {x y : α} (hx : x ∈ s) (hy : y ∈ s) : - edist (f x) (f y) ≤ eVariationOn f s := by - wlog hxy : y ≤ x generalizing x y - · rw [edist_comm] - exact this hy hx (le_of_not_le hxy) - let u : ℕ → α := fun n => if n = 0 then y else x - have hu : Monotone u := monotone_nat_of_le_succ fun - | 0 => hxy - | (_ + 1) => le_rfl - have us : ∀ i, u i ∈ s := fun - | 0 => hy - | (_ + 1) => hx - simpa only [Finset.sum_range_one] using sum_le f 1 hu us - -theorem eq_zero_iff (f : α → E) {s : Set α} : - eVariationOn f s = 0 ↔ ∀ x ∈ s, ∀ y ∈ s, edist (f x) (f y) = 0 := by - constructor - · rintro h x xs y ys - rw [← le_zero_iff, ← h] - exact edist_le f xs ys - · rintro h - dsimp only [eVariationOn] - rw [ENNReal.iSup_eq_zero] - rintro ⟨n, u, um, us⟩ - exact Finset.sum_eq_zero fun i _ => h _ (us i.succ) _ (us i) - -theorem constant_on {f : α → E} {s : Set α} (hf : (f '' s).Subsingleton) : - eVariationOn f s = 0 := by - rw [eq_zero_iff] - rintro x xs y ys - rw [hf ⟨x, xs, rfl⟩ ⟨y, ys, rfl⟩, edist_self] - -@[simp] -protected theorem subsingleton (f : α → E) {s : Set α} (hs : s.Subsingleton) : - eVariationOn f s = 0 := - constant_on (hs.image f) - -theorem lowerSemicontinuous_aux {ι : Type*} {F : ι → α → E} {p : Filter ι} {f : α → E} {s : Set α} - (Ffs : ∀ x ∈ s, Tendsto (fun i => F i x) p (𝓝 (f x))) {v : ℝ≥0∞} (hv : v < eVariationOn f s) : - ∀ᶠ n : ι in p, v < eVariationOn (F n) s := by - obtain ⟨⟨n, ⟨u, um, us⟩⟩, hlt⟩ : - ∃ p : ℕ × { u : ℕ → α // Monotone u ∧ ∀ i, u i ∈ s }, - v < ∑ i ∈ Finset.range p.1, edist (f ((p.2 : ℕ → α) (i + 1))) (f ((p.2 : ℕ → α) i)) := - lt_iSup_iff.mp hv - have : Tendsto (fun j => ∑ i ∈ Finset.range n, edist (F j (u (i + 1))) (F j (u i))) p - (𝓝 (∑ i ∈ Finset.range n, edist (f (u (i + 1))) (f (u i)))) := by - apply tendsto_finset_sum - exact fun i _ => Tendsto.edist (Ffs (u i.succ) (us i.succ)) (Ffs (u i) (us i)) - exact (this.eventually_const_lt hlt).mono fun i h => h.trans_le (sum_le (F i) n um us) - -/-- The map `(eVariationOn · s)` is lower semicontinuous for pointwise convergence *on `s`*. -Pointwise convergence on `s` is encoded here as uniform convergence on the family consisting of the -singletons of elements of `s`. --/ -protected theorem lowerSemicontinuous (s : Set α) : - LowerSemicontinuous fun f : α →ᵤ[s.image singleton] E => eVariationOn f s := fun f ↦ by - apply @lowerSemicontinuous_aux _ _ _ _ (UniformOnFun α E (s.image singleton)) id (𝓝 f) f s _ - simpa only [UniformOnFun.tendsto_iff_tendstoUniformlyOn, mem_image, forall_exists_index, and_imp, - forall_apply_eq_imp_iff₂, tendstoUniformlyOn_singleton_iff_tendsto] using @tendsto_id _ (𝓝 f) - -/-- The map `(eVariationOn · s)` is lower semicontinuous for uniform convergence on `s`. -/ -theorem lowerSemicontinuous_uniformOn (s : Set α) : - LowerSemicontinuous fun f : α →ᵤ[{s}] E => eVariationOn f s := fun f ↦ by - apply @lowerSemicontinuous_aux _ _ _ _ (UniformOnFun α E {s}) id (𝓝 f) f s _ - have := @tendsto_id _ (𝓝 f) - rw [UniformOnFun.tendsto_iff_tendstoUniformlyOn] at this - simp_rw [← tendstoUniformlyOn_singleton_iff_tendsto] - exact fun x xs => (this s rfl).mono (singleton_subset_iff.mpr xs) - -theorem _root_.BoundedVariationOn.dist_le {E : Type*} [PseudoMetricSpace E] {f : α → E} - {s : Set α} (h : BoundedVariationOn f s) {x y : α} (hx : x ∈ s) (hy : y ∈ s) : - dist (f x) (f y) ≤ (eVariationOn f s).toReal := by - rw [← ENNReal.ofReal_le_ofReal_iff ENNReal.toReal_nonneg, ENNReal.ofReal_toReal h, ← edist_dist] - exact edist_le f hx hy - -theorem _root_.BoundedVariationOn.sub_le {f : α → ℝ} {s : Set α} (h : BoundedVariationOn f s) - {x y : α} (hx : x ∈ s) (hy : y ∈ s) : f x - f y ≤ (eVariationOn f s).toReal := by - apply (le_abs_self _).trans - rw [← Real.dist_eq] - exact h.dist_le hx hy - -/-- Consider a monotone function `u` parameterizing some points of a set `s`. Given `x ∈ s`, then -one can find another monotone function `v` parameterizing the same points as `u`, with `x` added. -In particular, the variation of a function along `u` is bounded by its variation along `v`. -/ -theorem add_point (f : α → E) {s : Set α} {x : α} (hx : x ∈ s) (u : ℕ → α) (hu : Monotone u) - (us : ∀ i, u i ∈ s) (n : ℕ) : - ∃ (v : ℕ → α) (m : ℕ), Monotone v ∧ (∀ i, v i ∈ s) ∧ x ∈ v '' Iio m ∧ - (∑ i ∈ Finset.range n, edist (f (u (i + 1))) (f (u i))) ≤ - ∑ j ∈ Finset.range m, edist (f (v (j + 1))) (f (v j)) := by - rcases le_or_lt (u n) x with (h | h) - · let v i := if i ≤ n then u i else x - have vs : ∀ i, v i ∈ s := fun i ↦ by - simp only [v] - split_ifs - · exact us i - · exact hx - have hv : Monotone v := by - refine monotone_nat_of_le_succ fun i => ?_ - simp only [v] - rcases lt_trichotomy i n with (hi | rfl | hi) - · have : i + 1 ≤ n := Nat.succ_le_of_lt hi - simp only [hi.le, this, if_true] - exact hu (Nat.le_succ i) - · simp only [le_refl, if_true, add_le_iff_nonpos_right, Nat.le_zero, Nat.one_ne_zero, - if_false, h] - · have A : ¬i ≤ n := hi.not_le - have B : ¬i + 1 ≤ n := fun h => A (i.le_succ.trans h) - simp only [A, B, if_false, le_rfl] - refine ⟨v, n + 2, hv, vs, (mem_image _ _ _).2 ⟨n + 1, ?_, ?_⟩, ?_⟩ - · rw [mem_Iio]; exact Nat.lt_succ_self (n + 1) - · have : ¬n + 1 ≤ n := Nat.not_succ_le_self n - simp only [v, this, ite_eq_right_iff, IsEmpty.forall_iff] - · calc - (∑ i ∈ Finset.range n, edist (f (u (i + 1))) (f (u i))) = - ∑ i ∈ Finset.range n, edist (f (v (i + 1))) (f (v i)) := by - apply Finset.sum_congr rfl fun i hi => ?_ - simp only [Finset.mem_range] at hi - have : i + 1 ≤ n := Nat.succ_le_of_lt hi - simp only [v, hi.le, this, if_true] - _ ≤ ∑ j ∈ Finset.range (n + 2), edist (f (v (j + 1))) (f (v j)) := - Finset.sum_le_sum_of_subset (Finset.range_mono (Nat.le_add_right n 2)) - have exists_N : ∃ N, N ≤ n ∧ x < u N := ⟨n, le_rfl, h⟩ - let N := Nat.find exists_N - have hN : N ≤ n ∧ x < u N := Nat.find_spec exists_N - let w : ℕ → α := fun i => if i < N then u i else if i = N then x else u (i - 1) - have ws : ∀ i, w i ∈ s := by - dsimp only [w] - intro i - split_ifs - exacts [us _, hx, us _] - have hw : Monotone w := by - apply monotone_nat_of_le_succ fun i => ?_ - dsimp only [w] - rcases lt_trichotomy (i + 1) N with (hi | hi | hi) - · have : i < N := Nat.lt_of_le_of_lt (Nat.le_succ i) hi - simp only [hi, this, if_true] - exact hu (Nat.le_succ _) - · have A : i < N := hi ▸ i.lt_succ_self - have B : ¬i + 1 < N := by rw [← hi]; exact fun h => h.ne rfl - rw [if_pos A, if_neg B, if_pos hi] - have T := Nat.find_min exists_N A - push_neg at T - exact T (A.le.trans hN.1) - · have A : ¬i < N := (Nat.lt_succ_iff.mp hi).not_lt - have B : ¬i + 1 < N := hi.not_lt - have C : ¬i + 1 = N := hi.ne.symm - have D : i + 1 - 1 = i := Nat.pred_succ i - rw [if_neg A, if_neg B, if_neg C, D] - split_ifs - · exact hN.2.le.trans (hu (le_of_not_lt A)) - · exact hu (Nat.pred_le _) - refine ⟨w, n + 1, hw, ws, (mem_image _ _ _).2 ⟨N, hN.1.trans_lt (Nat.lt_succ_self n), ?_⟩, ?_⟩ - · dsimp only [w]; rw [if_neg (lt_irrefl N), if_pos rfl] - rcases eq_or_lt_of_le (zero_le N) with (Npos | Npos) - · calc - (∑ i ∈ Finset.range n, edist (f (u (i + 1))) (f (u i))) = - ∑ i ∈ Finset.range n, edist (f (w (1 + i + 1))) (f (w (1 + i))) := by - apply Finset.sum_congr rfl fun i _hi => ?_ - dsimp only [w] - simp only [← Npos, Nat.not_lt_zero, Nat.add_succ_sub_one, add_zero, if_false, - add_eq_zero, Nat.one_ne_zero, false_and, Nat.succ_add_sub_one, zero_add] - rw [add_comm 1 i] - _ = ∑ i ∈ Finset.Ico 1 (n + 1), edist (f (w (i + 1))) (f (w i)) := by - rw [Finset.range_eq_Ico] - exact Finset.sum_Ico_add (fun i => edist (f (w (i + 1))) (f (w i))) 0 n 1 - _ ≤ ∑ j ∈ Finset.range (n + 1), edist (f (w (j + 1))) (f (w j)) := by - apply Finset.sum_le_sum_of_subset _ - rw [Finset.range_eq_Ico] - exact Finset.Ico_subset_Ico zero_le_one le_rfl - · calc - (∑ i ∈ Finset.range n, edist (f (u (i + 1))) (f (u i))) = - ((∑ i ∈ Finset.Ico 0 (N - 1), edist (f (u (i + 1))) (f (u i))) + - ∑ i ∈ Finset.Ico (N - 1) N, edist (f (u (i + 1))) (f (u i))) + - ∑ i ∈ Finset.Ico N n, edist (f (u (i + 1))) (f (u i)) := by - rw [Finset.sum_Ico_consecutive, Finset.sum_Ico_consecutive, Finset.range_eq_Ico] - · exact zero_le _ - · exact hN.1 - · exact zero_le _ - · exact Nat.pred_le _ - _ = (∑ i ∈ Finset.Ico 0 (N - 1), edist (f (w (i + 1))) (f (w i))) + - edist (f (u N)) (f (u (N - 1))) + - ∑ i ∈ Finset.Ico N n, edist (f (w (1 + i + 1))) (f (w (1 + i))) := by - congr 1 - · congr 1 - · apply Finset.sum_congr rfl fun i hi => ?_ - simp only [Finset.mem_Ico, zero_le', true_and] at hi - dsimp only [w] - have A : i + 1 < N := Nat.lt_pred_iff.1 hi - have B : i < N := Nat.lt_of_succ_lt A - rw [if_pos A, if_pos B] - · have A : N - 1 + 1 = N := Nat.succ_pred_eq_of_pos Npos - have : Finset.Ico (N - 1) N = {N - 1} := by rw [← Nat.Ico_succ_singleton, A] - simp only [this, A, Finset.sum_singleton] - · apply Finset.sum_congr rfl fun i hi => ?_ - rw [Finset.mem_Ico] at hi - dsimp only [w] - have A : ¬1 + i + 1 < N := by omega - have B : ¬1 + i + 1 = N := by omega - have C : ¬1 + i < N := by omega - have D : ¬1 + i = N := by omega - rw [if_neg A, if_neg B, if_neg C, if_neg D] - congr 3 <;> · rw [add_comm, Nat.sub_one]; apply Nat.pred_succ - _ = (∑ i ∈ Finset.Ico 0 (N - 1), edist (f (w (i + 1))) (f (w i))) + - edist (f (w (N + 1))) (f (w (N - 1))) + - ∑ i ∈ Finset.Ico (N + 1) (n + 1), edist (f (w (i + 1))) (f (w i)) := by - congr 1 - · congr 1 - · dsimp only [w] - have A : ¬N + 1 < N := Nat.not_succ_lt_self - have B : N - 1 < N := Nat.pred_lt Npos.ne' - simp only [A, not_and, not_lt, Nat.succ_ne_self, Nat.add_succ_sub_one, add_zero, - if_false, B, if_true] - · exact Finset.sum_Ico_add (fun i => edist (f (w (i + 1))) (f (w i))) N n 1 - _ ≤ ((∑ i ∈ Finset.Ico 0 (N - 1), edist (f (w (i + 1))) (f (w i))) + - ∑ i ∈ Finset.Ico (N - 1) (N + 1), edist (f (w (i + 1))) (f (w i))) + - ∑ i ∈ Finset.Ico (N + 1) (n + 1), edist (f (w (i + 1))) (f (w i)) := by - refine add_le_add (add_le_add le_rfl ?_) le_rfl - have A : N - 1 + 1 = N := Nat.succ_pred_eq_of_pos Npos - have B : N - 1 + 1 < N + 1 := A.symm ▸ N.lt_succ_self - have C : N - 1 < N + 1 := lt_of_le_of_lt N.pred_le N.lt_succ_self - rw [Finset.sum_eq_sum_Ico_succ_bot C, Finset.sum_eq_sum_Ico_succ_bot B, A, Finset.Ico_self, - Finset.sum_empty, add_zero, add_comm (edist _ _)] - exact edist_triangle _ _ _ - _ = ∑ j ∈ Finset.range (n + 1), edist (f (w (j + 1))) (f (w j)) := by - rw [Finset.sum_Ico_consecutive, Finset.sum_Ico_consecutive, Finset.range_eq_Ico] - · exact zero_le _ - · exact Nat.succ_le_succ hN.left - · exact zero_le _ - · exact N.pred_le.trans N.le_succ - -/-- The variation of a function on the union of two sets `s` and `t`, with `s` to the left of `t`, -bounds the sum of the variations along `s` and `t`. -/ -theorem add_le_union (f : α → E) {s t : Set α} (h : ∀ x ∈ s, ∀ y ∈ t, x ≤ y) : - eVariationOn f s + eVariationOn f t ≤ eVariationOn f (s ∪ t) := by - by_cases hs : s = ∅ - · simp [hs] - have : Nonempty { u // Monotone u ∧ ∀ i : ℕ, u i ∈ s } := - nonempty_monotone_mem (nonempty_iff_ne_empty.2 hs) - by_cases ht : t = ∅ - · simp [ht] - have : Nonempty { u // Monotone u ∧ ∀ i : ℕ, u i ∈ t } := - nonempty_monotone_mem (nonempty_iff_ne_empty.2 ht) - refine ENNReal.iSup_add_iSup_le ?_ - /- We start from two sequences `u` and `v` along `s` and `t` respectively, and we build a new - sequence `w` along `s ∪ t` by juxtaposing them. Its variation is larger than the sum of the - variations. -/ - rintro ⟨n, ⟨u, hu, us⟩⟩ ⟨m, ⟨v, hv, vt⟩⟩ - let w i := if i ≤ n then u i else v (i - (n + 1)) - have wst : ∀ i, w i ∈ s ∪ t := by - intro i - by_cases hi : i ≤ n - · simp [w, hi, us] - · simp [w, hi, vt] - have hw : Monotone w := by - intro i j hij - dsimp only [w] - split_ifs with h_1 h_2 h_2 - · exact hu hij - · apply h _ (us _) _ (vt _) - · exfalso; exact h_1 (hij.trans h_2) - · apply hv (tsub_le_tsub hij le_rfl) - calc - ((∑ i ∈ Finset.range n, edist (f (u (i + 1))) (f (u i))) + - ∑ i ∈ Finset.range m, edist (f (v (i + 1))) (f (v i))) = - (∑ i ∈ Finset.range n, edist (f (w (i + 1))) (f (w i))) + - ∑ i ∈ Finset.range m, edist (f (w (n + 1 + i + 1))) (f (w (n + 1 + i))) := by - dsimp only [w] - congr 1 - · refine Finset.sum_congr rfl fun i hi => ?_ - simp only [Finset.mem_range] at hi - have : i + 1 ≤ n := Nat.succ_le_of_lt hi - simp [hi.le, this] - · refine Finset.sum_congr rfl fun i hi => ?_ - simp only [Finset.mem_range] at hi - have B : ¬n + 1 + i ≤ n := by omega - have A : ¬n + 1 + i + 1 ≤ n := fun h => B ((n + 1 + i).le_succ.trans h) - have C : n + 1 + i - n = i + 1 := by - rw [tsub_eq_iff_eq_add_of_le] - · abel - · exact n.le_succ.trans (n.succ.le_add_right i) - simp only [A, B, C, Nat.succ_sub_succ_eq_sub, if_false, add_tsub_cancel_left] - _ = (∑ i ∈ Finset.range n, edist (f (w (i + 1))) (f (w i))) + - ∑ i ∈ Finset.Ico (n + 1) (n + 1 + m), edist (f (w (i + 1))) (f (w i)) := by - congr 1 - rw [Finset.range_eq_Ico] - convert Finset.sum_Ico_add (fun i : ℕ => edist (f (w (i + 1))) (f (w i))) 0 m (n + 1) - using 3 <;> abel - _ ≤ ∑ i ∈ Finset.range (n + 1 + m), edist (f (w (i + 1))) (f (w i)) := by - rw [← Finset.sum_union] - · apply Finset.sum_le_sum_of_subset _ - rintro i hi - simp only [Finset.mem_union, Finset.mem_range, Finset.mem_Ico] at hi ⊢ - cases' hi with hi hi - · exact lt_of_lt_of_le hi (n.le_succ.trans (n.succ.le_add_right m)) - · exact hi.2 - · refine Finset.disjoint_left.2 fun i hi h'i => ?_ - simp only [Finset.mem_Ico, Finset.mem_range] at hi h'i - exact hi.not_lt (Nat.lt_of_succ_le h'i.left) - _ ≤ eVariationOn f (s ∪ t) := sum_le f _ hw wst - -/-- If a set `s` is to the left of a set `t`, and both contain the boundary point `x`, then -the variation of `f` along `s ∪ t` is the sum of the variations. -/ -theorem union (f : α → E) {s t : Set α} {x : α} (hs : IsGreatest s x) (ht : IsLeast t x) : - eVariationOn f (s ∪ t) = eVariationOn f s + eVariationOn f t := by - classical - apply le_antisymm _ (eVariationOn.add_le_union f fun a ha b hb => le_trans (hs.2 ha) (ht.2 hb)) - apply iSup_le _ - rintro ⟨n, ⟨u, hu, ust⟩⟩ - obtain ⟨v, m, hv, vst, xv, huv⟩ : ∃ (v : ℕ → α) (m : ℕ), - Monotone v ∧ (∀ i, v i ∈ s ∪ t) ∧ x ∈ v '' Iio m ∧ - (∑ i ∈ Finset.range n, edist (f (u (i + 1))) (f (u i))) ≤ - ∑ j ∈ Finset.range m, edist (f (v (j + 1))) (f (v j)) := - eVariationOn.add_point f (mem_union_left t hs.1) u hu ust n - obtain ⟨N, hN, Nx⟩ : ∃ N, N < m ∧ v N = x := xv - calc - (∑ j ∈ Finset.range n, edist (f (u (j + 1))) (f (u j))) ≤ - ∑ j ∈ Finset.range m, edist (f (v (j + 1))) (f (v j)) := - huv - _ = (∑ j ∈ Finset.Ico 0 N, edist (f (v (j + 1))) (f (v j))) + - ∑ j ∈ Finset.Ico N m, edist (f (v (j + 1))) (f (v j)) := by - rw [Finset.range_eq_Ico, Finset.sum_Ico_consecutive _ (zero_le _) hN.le] - _ ≤ eVariationOn f s + eVariationOn f t := by - refine add_le_add ?_ ?_ - · apply sum_le_of_monotoneOn_Icc _ (hv.monotoneOn _) fun i hi => ?_ - rcases vst i with (h | h); · exact h - have : v i = x := by - apply le_antisymm - · rw [← Nx]; exact hv hi.2 - · exact ht.2 h - rw [this] - exact hs.1 - · apply sum_le_of_monotoneOn_Icc _ (hv.monotoneOn _) fun i hi => ?_ - rcases vst i with (h | h); swap; · exact h - have : v i = x := by - apply le_antisymm - · exact hs.2 h - · rw [← Nx]; exact hv hi.1 - rw [this] - exact ht.1 - -theorem Icc_add_Icc (f : α → E) {s : Set α} {a b c : α} (hab : a ≤ b) (hbc : b ≤ c) (hb : b ∈ s) : - eVariationOn f (s ∩ Icc a b) + eVariationOn f (s ∩ Icc b c) = eVariationOn f (s ∩ Icc a c) := by - have A : IsGreatest (s ∩ Icc a b) b := - ⟨⟨hb, hab, le_rfl⟩, inter_subset_right.trans Icc_subset_Iic_self⟩ - have B : IsLeast (s ∩ Icc b c) b := - ⟨⟨hb, le_rfl, hbc⟩, inter_subset_right.trans Icc_subset_Ici_self⟩ - rw [← eVariationOn.union f A B, ← inter_union_distrib_left, Icc_union_Icc_eq_Icc hab hbc] - -section Monotone - -variable {β : Type*} [LinearOrder β] - -theorem comp_le_of_monotoneOn (f : α → E) {s : Set α} {t : Set β} (φ : β → α) (hφ : MonotoneOn φ t) - (φst : MapsTo φ t s) : eVariationOn (f ∘ φ) t ≤ eVariationOn f s := - iSup_le fun ⟨n, u, hu, ut⟩ => - le_iSup_of_le ⟨n, φ ∘ u, fun x y xy => hφ (ut x) (ut y) (hu xy), fun i => φst (ut i)⟩ le_rfl - -theorem comp_le_of_antitoneOn (f : α → E) {s : Set α} {t : Set β} (φ : β → α) (hφ : AntitoneOn φ t) - (φst : MapsTo φ t s) : eVariationOn (f ∘ φ) t ≤ eVariationOn f s := by - refine iSup_le ?_ - rintro ⟨n, u, hu, ut⟩ - rw [← Finset.sum_range_reflect] - refine (Finset.sum_congr rfl fun x hx => ?_).trans_le <| le_iSup_of_le - ⟨n, fun i => φ (u <| n - i), fun x y xy => hφ (ut _) (ut _) (hu <| Nat.sub_le_sub_left xy n), - fun i => φst (ut _)⟩ - le_rfl - rw [Finset.mem_range] at hx - dsimp only [Subtype.coe_mk, Function.comp_apply] - rw [edist_comm] - congr 4 <;> omega - -theorem comp_eq_of_monotoneOn (f : α → E) {t : Set β} (φ : β → α) (hφ : MonotoneOn φ t) : - eVariationOn (f ∘ φ) t = eVariationOn f (φ '' t) := by - apply le_antisymm (comp_le_of_monotoneOn f φ hφ (mapsTo_image φ t)) - cases isEmpty_or_nonempty β - · convert zero_le (_ : ℝ≥0∞) - exact eVariationOn.subsingleton f <| - (subsingleton_of_subsingleton.image _).anti (surjOn_image φ t) - let ψ := φ.invFunOn t - have ψφs : EqOn (φ ∘ ψ) id (φ '' t) := (surjOn_image φ t).rightInvOn_invFunOn - have ψts : MapsTo ψ (φ '' t) t := (surjOn_image φ t).mapsTo_invFunOn - have hψ : MonotoneOn ψ (φ '' t) := Function.monotoneOn_of_rightInvOn_of_mapsTo hφ ψφs ψts - change eVariationOn (f ∘ id) (φ '' t) ≤ eVariationOn (f ∘ φ) t - rw [← eq_of_eqOn (ψφs.comp_left : EqOn (f ∘ φ ∘ ψ) (f ∘ id) (φ '' t))] - exact comp_le_of_monotoneOn _ ψ hψ ψts - -theorem comp_inter_Icc_eq_of_monotoneOn (f : α → E) {t : Set β} (φ : β → α) (hφ : MonotoneOn φ t) - {x y : β} (hx : x ∈ t) (hy : y ∈ t) : - eVariationOn (f ∘ φ) (t ∩ Icc x y) = eVariationOn f (φ '' t ∩ Icc (φ x) (φ y)) := by - rcases le_total x y with (h | h) - · convert comp_eq_of_monotoneOn f φ (hφ.mono Set.inter_subset_left) - apply le_antisymm - · rintro _ ⟨⟨u, us, rfl⟩, vφx, vφy⟩ - rcases le_total x u with (xu | ux) - · rcases le_total u y with (uy | yu) - · exact ⟨u, ⟨us, ⟨xu, uy⟩⟩, rfl⟩ - · rw [le_antisymm vφy (hφ hy us yu)] - exact ⟨y, ⟨hy, ⟨h, le_rfl⟩⟩, rfl⟩ - · rw [← le_antisymm vφx (hφ us hx ux)] - exact ⟨x, ⟨hx, ⟨le_rfl, h⟩⟩, rfl⟩ - · rintro _ ⟨u, ⟨⟨hu, xu, uy⟩, rfl⟩⟩ - exact ⟨⟨u, hu, rfl⟩, ⟨hφ hx hu xu, hφ hu hy uy⟩⟩ - · rw [eVariationOn.subsingleton, eVariationOn.subsingleton] - exacts [(Set.subsingleton_Icc_of_ge (hφ hy hx h)).anti Set.inter_subset_right, - (Set.subsingleton_Icc_of_ge h).anti Set.inter_subset_right] - -theorem comp_eq_of_antitoneOn (f : α → E) {t : Set β} (φ : β → α) (hφ : AntitoneOn φ t) : - eVariationOn (f ∘ φ) t = eVariationOn f (φ '' t) := by - apply le_antisymm (comp_le_of_antitoneOn f φ hφ (mapsTo_image φ t)) - cases isEmpty_or_nonempty β - · convert zero_le (_ : ℝ≥0∞) - exact eVariationOn.subsingleton f <| (subsingleton_of_subsingleton.image _).anti - (surjOn_image φ t) - let ψ := φ.invFunOn t - have ψφs : EqOn (φ ∘ ψ) id (φ '' t) := (surjOn_image φ t).rightInvOn_invFunOn - have ψts := (surjOn_image φ t).mapsTo_invFunOn - have hψ : AntitoneOn ψ (φ '' t) := Function.antitoneOn_of_rightInvOn_of_mapsTo hφ ψφs ψts - change eVariationOn (f ∘ id) (φ '' t) ≤ eVariationOn (f ∘ φ) t - rw [← eq_of_eqOn (ψφs.comp_left : EqOn (f ∘ φ ∘ ψ) (f ∘ id) (φ '' t))] - exact comp_le_of_antitoneOn _ ψ hψ ψts - -open OrderDual - -theorem comp_ofDual (f : α → E) (s : Set α) : - eVariationOn (f ∘ ofDual) (ofDual ⁻¹' s) = eVariationOn f s := by - convert comp_eq_of_antitoneOn f ofDual fun _ _ _ _ => id - simp only [Equiv.image_preimage] - -end Monotone - -end eVariationOn - -/-! ## Monotone functions and bounded variation -/ - -theorem MonotoneOn.eVariationOn_le {f : α → ℝ} {s : Set α} (hf : MonotoneOn f s) {a b : α} - (as : a ∈ s) (bs : b ∈ s) : eVariationOn f (s ∩ Icc a b) ≤ ENNReal.ofReal (f b - f a) := by - apply iSup_le _ - rintro ⟨n, ⟨u, hu, us⟩⟩ - calc - (∑ i ∈ Finset.range n, edist (f (u (i + 1))) (f (u i))) = - ∑ i ∈ Finset.range n, ENNReal.ofReal (f (u (i + 1)) - f (u i)) := by - refine Finset.sum_congr rfl fun i hi => ?_ - simp only [Finset.mem_range] at hi - rw [edist_dist, Real.dist_eq, abs_of_nonneg] - exact sub_nonneg_of_le (hf (us i).1 (us (i + 1)).1 (hu (Nat.le_succ _))) - _ = ENNReal.ofReal (∑ i ∈ Finset.range n, (f (u (i + 1)) - f (u i))) := by - rw [ENNReal.ofReal_sum_of_nonneg] - intro i _ - exact sub_nonneg_of_le (hf (us i).1 (us (i + 1)).1 (hu (Nat.le_succ _))) - _ = ENNReal.ofReal (f (u n) - f (u 0)) := by rw [Finset.sum_range_sub fun i => f (u i)] - _ ≤ ENNReal.ofReal (f b - f a) := by - apply ENNReal.ofReal_le_ofReal - exact sub_le_sub (hf (us n).1 bs (us n).2.2) (hf as (us 0).1 (us 0).2.1) - -theorem MonotoneOn.locallyBoundedVariationOn {f : α → ℝ} {s : Set α} (hf : MonotoneOn f s) : - LocallyBoundedVariationOn f s := fun _ _ as bs => - ((hf.eVariationOn_le as bs).trans_lt ENNReal.ofReal_lt_top).ne - -/-- The **signed** variation of `f` on the interval `Icc a b` intersected with the set `s`, -squashed to a real (therefore only really meaningful if the variation is finite) --/ -noncomputable def variationOnFromTo (f : α → E) (s : Set α) (a b : α) : ℝ := - if a ≤ b then (eVariationOn f (s ∩ Icc a b)).toReal else -(eVariationOn f (s ∩ Icc b a)).toReal - -namespace variationOnFromTo - -variable (f : α → E) (s : Set α) - -protected theorem self (a : α) : variationOnFromTo f s a a = 0 := by - dsimp only [variationOnFromTo] - rw [if_pos le_rfl, Icc_self, eVariationOn.subsingleton, ENNReal.zero_toReal] - exact fun x hx y hy => hx.2.trans hy.2.symm - -protected theorem nonneg_of_le {a b : α} (h : a ≤ b) : 0 ≤ variationOnFromTo f s a b := by - simp only [variationOnFromTo, if_pos h, ENNReal.toReal_nonneg] - -protected theorem eq_neg_swap (a b : α) : - variationOnFromTo f s a b = -variationOnFromTo f s b a := by - rcases lt_trichotomy a b with (ab | rfl | ba) - · simp only [variationOnFromTo, if_pos ab.le, if_neg ab.not_le, neg_neg] - · simp only [variationOnFromTo.self, neg_zero] - · simp only [variationOnFromTo, if_pos ba.le, if_neg ba.not_le, neg_neg] - -protected theorem nonpos_of_ge {a b : α} (h : b ≤ a) : variationOnFromTo f s a b ≤ 0 := by - rw [variationOnFromTo.eq_neg_swap] - exact neg_nonpos_of_nonneg (variationOnFromTo.nonneg_of_le f s h) - -protected theorem eq_of_le {a b : α} (h : a ≤ b) : - variationOnFromTo f s a b = (eVariationOn f (s ∩ Icc a b)).toReal := - if_pos h - -protected theorem eq_of_ge {a b : α} (h : b ≤ a) : - variationOnFromTo f s a b = -(eVariationOn f (s ∩ Icc b a)).toReal := by - rw [variationOnFromTo.eq_neg_swap, neg_inj, variationOnFromTo.eq_of_le f s h] - -protected theorem add {f : α → E} {s : Set α} (hf : LocallyBoundedVariationOn f s) {a b c : α} - (ha : a ∈ s) (hb : b ∈ s) (hc : c ∈ s) : - variationOnFromTo f s a b + variationOnFromTo f s b c = variationOnFromTo f s a c := by - symm - refine additive_of_isTotal ((· : α) ≤ ·) (variationOnFromTo f s) (· ∈ s) ?_ ?_ ha hb hc - · rintro x y _xs _ys - simp only [variationOnFromTo.eq_neg_swap f s y x, Subtype.coe_mk, add_neg_cancel, - forall_true_left] - · rintro x y z xy yz xs ys zs - rw [variationOnFromTo.eq_of_le f s xy, variationOnFromTo.eq_of_le f s yz, - variationOnFromTo.eq_of_le f s (xy.trans yz), - ← ENNReal.toReal_add (hf x y xs ys) (hf y z ys zs), eVariationOn.Icc_add_Icc f xy yz ys] - -variable {f s} in -protected theorem edist_zero_of_eq_zero (hf : LocallyBoundedVariationOn f s) - {a b : α} (ha : a ∈ s) (hb : b ∈ s) (h : variationOnFromTo f s a b = 0) : - edist (f a) (f b) = 0 := by - wlog h' : a ≤ b - · rw [edist_comm] - apply this hf hb ha _ (le_of_not_le h') - rw [variationOnFromTo.eq_neg_swap, h, neg_zero] - · apply le_antisymm _ (zero_le _) - rw [← ENNReal.ofReal_zero, ← h, variationOnFromTo.eq_of_le f s h', - ENNReal.ofReal_toReal (hf a b ha hb)] - apply eVariationOn.edist_le - exacts [⟨ha, ⟨le_rfl, h'⟩⟩, ⟨hb, ⟨h', le_rfl⟩⟩] - -protected theorem eq_left_iff {f : α → E} {s : Set α} (hf : LocallyBoundedVariationOn f s) - {a b c : α} (ha : a ∈ s) (hb : b ∈ s) (hc : c ∈ s) : - variationOnFromTo f s a b = variationOnFromTo f s a c ↔ variationOnFromTo f s b c = 0 := by - simp only [← variationOnFromTo.add hf ha hb hc, self_eq_add_right] - -protected theorem eq_zero_iff_of_le {f : α → E} {s : Set α} (hf : LocallyBoundedVariationOn f s) - {a b : α} (ha : a ∈ s) (hb : b ∈ s) (ab : a ≤ b) : - variationOnFromTo f s a b = 0 ↔ - ∀ ⦃x⦄ (_hx : x ∈ s ∩ Icc a b) ⦃y⦄ (_hy : y ∈ s ∩ Icc a b), edist (f x) (f y) = 0 := by - rw [variationOnFromTo.eq_of_le _ _ ab, ENNReal.toReal_eq_zero_iff, or_iff_left (hf a b ha hb), - eVariationOn.eq_zero_iff] - -protected theorem eq_zero_iff_of_ge {f : α → E} {s : Set α} (hf : LocallyBoundedVariationOn f s) - {a b : α} (ha : a ∈ s) (hb : b ∈ s) (ba : b ≤ a) : - variationOnFromTo f s a b = 0 ↔ - ∀ ⦃x⦄ (_hx : x ∈ s ∩ Icc b a) ⦃y⦄ (_hy : y ∈ s ∩ Icc b a), edist (f x) (f y) = 0 := by - rw [variationOnFromTo.eq_of_ge _ _ ba, neg_eq_zero, ENNReal.toReal_eq_zero_iff, - or_iff_left (hf b a hb ha), eVariationOn.eq_zero_iff] - -protected theorem eq_zero_iff {f : α → E} {s : Set α} (hf : LocallyBoundedVariationOn f s) {a b : α} - (ha : a ∈ s) (hb : b ∈ s) : - variationOnFromTo f s a b = 0 ↔ - ∀ ⦃x⦄ (_hx : x ∈ s ∩ uIcc a b) ⦃y⦄ (_hy : y ∈ s ∩ uIcc a b), edist (f x) (f y) = 0 := by - rcases le_total a b with (ab | ba) - · rw [uIcc_of_le ab] - exact variationOnFromTo.eq_zero_iff_of_le hf ha hb ab - · rw [uIcc_of_ge ba] - exact variationOnFromTo.eq_zero_iff_of_ge hf ha hb ba - -variable {f} {s} - -protected theorem monotoneOn (hf : LocallyBoundedVariationOn f s) {a : α} (as : a ∈ s) : - MonotoneOn (variationOnFromTo f s a) s := by - rintro b bs c cs bc - rw [← variationOnFromTo.add hf as bs cs] - exact le_add_of_nonneg_right (variationOnFromTo.nonneg_of_le f s bc) - -protected theorem antitoneOn (hf : LocallyBoundedVariationOn f s) {b : α} (bs : b ∈ s) : - AntitoneOn (fun a => variationOnFromTo f s a b) s := by - rintro a as c cs ac - dsimp only - rw [← variationOnFromTo.add hf as cs bs] - exact le_add_of_nonneg_left (variationOnFromTo.nonneg_of_le f s ac) - -protected theorem sub_self_monotoneOn {f : α → ℝ} {s : Set α} (hf : LocallyBoundedVariationOn f s) - {a : α} (as : a ∈ s) : MonotoneOn (variationOnFromTo f s a - f) s := by - rintro b bs c cs bc - rw [Pi.sub_apply, Pi.sub_apply, le_sub_iff_add_le, add_comm_sub, ← le_sub_iff_add_le'] - calc - f c - f b ≤ |f c - f b| := le_abs_self _ - _ = dist (f b) (f c) := by rw [dist_comm, Real.dist_eq] - _ ≤ variationOnFromTo f s b c := by - rw [variationOnFromTo.eq_of_le f s bc, dist_edist] - apply ENNReal.toReal_mono (hf b c bs cs) - apply eVariationOn.edist_le f - exacts [⟨bs, le_rfl, bc⟩, ⟨cs, bc, le_rfl⟩] - _ = variationOnFromTo f s a c - variationOnFromTo f s a b := by - rw [← variationOnFromTo.add hf as bs cs, add_sub_cancel_left] - -protected theorem comp_eq_of_monotoneOn {β : Type*} [LinearOrder β] (f : α → E) {t : Set β} - (φ : β → α) (hφ : MonotoneOn φ t) {x y : β} (hx : x ∈ t) (hy : y ∈ t) : - variationOnFromTo (f ∘ φ) t x y = variationOnFromTo f (φ '' t) (φ x) (φ y) := by - rcases le_total x y with (h | h) - · rw [variationOnFromTo.eq_of_le _ _ h, variationOnFromTo.eq_of_le _ _ (hφ hx hy h), - eVariationOn.comp_inter_Icc_eq_of_monotoneOn f φ hφ hx hy] - · rw [variationOnFromTo.eq_of_ge _ _ h, variationOnFromTo.eq_of_ge _ _ (hφ hy hx h), - eVariationOn.comp_inter_Icc_eq_of_monotoneOn f φ hφ hy hx] - -end variationOnFromTo - -/-- If a real valued function has bounded variation on a set, then it is a difference of monotone -functions there. -/ -theorem LocallyBoundedVariationOn.exists_monotoneOn_sub_monotoneOn {f : α → ℝ} {s : Set α} - (h : LocallyBoundedVariationOn f s) : - ∃ p q : α → ℝ, MonotoneOn p s ∧ MonotoneOn q s ∧ f = p - q := by - rcases eq_empty_or_nonempty s with (rfl | ⟨c, cs⟩) - · exact ⟨f, 0, subsingleton_empty.monotoneOn _, subsingleton_empty.monotoneOn _, - (sub_zero f).symm⟩ - · exact ⟨_, _, variationOnFromTo.monotoneOn h cs, variationOnFromTo.sub_self_monotoneOn h cs, - (sub_sub_cancel _ _).symm⟩ - -/-! ## Lipschitz functions and bounded variation -/ - -section LipschitzOnWith - -variable {F : Type*} [PseudoEMetricSpace F] - -theorem LipschitzOnWith.comp_eVariationOn_le {f : E → F} {C : ℝ≥0} {t : Set E} - (h : LipschitzOnWith C f t) {g : α → E} {s : Set α} (hg : MapsTo g s t) : - eVariationOn (f ∘ g) s ≤ C * eVariationOn g s := by - apply iSup_le _ - rintro ⟨n, ⟨u, hu, us⟩⟩ - calc - (∑ i ∈ Finset.range n, edist (f (g (u (i + 1)))) (f (g (u i)))) ≤ - ∑ i ∈ Finset.range n, C * edist (g (u (i + 1))) (g (u i)) := - Finset.sum_le_sum fun i _ => h (hg (us _)) (hg (us _)) - _ = C * ∑ i ∈ Finset.range n, edist (g (u (i + 1))) (g (u i)) := by rw [Finset.mul_sum] - _ ≤ C * eVariationOn g s := mul_le_mul_left' (eVariationOn.sum_le _ _ hu us) _ - -theorem LipschitzOnWith.comp_boundedVariationOn {f : E → F} {C : ℝ≥0} {t : Set E} - (hf : LipschitzOnWith C f t) {g : α → E} {s : Set α} (hg : MapsTo g s t) - (h : BoundedVariationOn g s) : BoundedVariationOn (f ∘ g) s := - ne_top_of_le_ne_top (ENNReal.mul_ne_top ENNReal.coe_ne_top h) (hf.comp_eVariationOn_le hg) - -theorem LipschitzOnWith.comp_locallyBoundedVariationOn {f : E → F} {C : ℝ≥0} {t : Set E} - (hf : LipschitzOnWith C f t) {g : α → E} {s : Set α} (hg : MapsTo g s t) - (h : LocallyBoundedVariationOn g s) : LocallyBoundedVariationOn (f ∘ g) s := - fun x y xs ys => - hf.comp_boundedVariationOn (hg.mono_left inter_subset_left) (h x y xs ys) - -theorem LipschitzWith.comp_boundedVariationOn {f : E → F} {C : ℝ≥0} (hf : LipschitzWith C f) - {g : α → E} {s : Set α} (h : BoundedVariationOn g s) : BoundedVariationOn (f ∘ g) s := - hf.lipschitzOnWith.comp_boundedVariationOn (mapsTo_univ _ _) h - -theorem LipschitzWith.comp_locallyBoundedVariationOn {f : E → F} {C : ℝ≥0} - (hf : LipschitzWith C f) {g : α → E} {s : Set α} (h : LocallyBoundedVariationOn g s) : - LocallyBoundedVariationOn (f ∘ g) s := - hf.lipschitzOnWith.comp_locallyBoundedVariationOn (mapsTo_univ _ _) h - -theorem LipschitzOnWith.locallyBoundedVariationOn {f : ℝ → E} {C : ℝ≥0} {s : Set ℝ} - (hf : LipschitzOnWith C f s) : LocallyBoundedVariationOn f s := - hf.comp_locallyBoundedVariationOn (mapsTo_id _) - (@monotoneOn_id ℝ _ s).locallyBoundedVariationOn - -theorem LipschitzWith.locallyBoundedVariationOn {f : ℝ → E} {C : ℝ≥0} (hf : LipschitzWith C f) - (s : Set ℝ) : LocallyBoundedVariationOn f s := - hf.lipschitzOnWith.locallyBoundedVariationOn - -end LipschitzOnWith - -/-! ## Almost everywhere differentiability of functions with locally bounded variation -/ +/-! ## -/ variable {V : Type*} [NormedAddCommGroup V] [NormedSpace ℝ V] [FiniteDimensional ℝ V] diff --git a/Mathlib/Analysis/ConstantSpeed.lean b/Mathlib/Analysis/ConstantSpeed.lean index 268a8fb396472..308cfdf9f15ce 100644 --- a/Mathlib/Analysis/ConstantSpeed.lean +++ b/Mathlib/Analysis/ConstantSpeed.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémi Bottinelli -/ import Mathlib.Data.Set.Function -import Mathlib.Analysis.BoundedVariation +import Mathlib.Analysis.RCLike.Basic +import Mathlib.Topology.EMetricSpace.BoundedVariation /-! # Constant speed @@ -40,7 +41,7 @@ arc-length, parameterization open scoped NNReal ENNReal -open Set MeasureTheory +open Set variable {α : Type*} [LinearOrder α] {E : Type*} [PseudoEMetricSpace E] variable (f : ℝ → E) (s : Set ℝ) (l : ℝ≥0) diff --git a/Mathlib/Topology/EMetricSpace/BoundedVariation.lean b/Mathlib/Topology/EMetricSpace/BoundedVariation.lean new file mode 100644 index 0000000000000..a8258f7c11588 --- /dev/null +++ b/Mathlib/Topology/EMetricSpace/BoundedVariation.lean @@ -0,0 +1,777 @@ +/- +Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel +-/ +import Mathlib.Order.Interval.Set.ProjIcc +import Mathlib.Topology.Semicontinuous +import Mathlib.Topology.UniformSpace.UniformConvergenceTopology + +/-! +# Functions of bounded variation + +We study functions of bounded variation. In particular, we show that a bounded variation function +is a difference of monotone functions, and differentiable almost everywhere. This implies that +Lipschitz functions from the real line into finite-dimensional vector space are also differentiable +almost everywhere. + +## Main definitions and results + +* `eVariationOn f s` is the total variation of the function `f` on the set `s`, in `ℝ≥0∞`. +* `BoundedVariationOn f s` registers that the variation of `f` on `s` is finite. +* `LocallyBoundedVariationOn f s` registers that `f` has finite variation on any compact + subinterval of `s`. +* `variationOnFromTo f s a b` is the signed variation of `f` on `s ∩ Icc a b`, converted to `ℝ`. + +* `eVariationOn.Icc_add_Icc` states that the variation of `f` on `[a, c]` is the sum of its + variations on `[a, b]` and `[b, c]`. +* `LocallyBoundedVariationOn.exists_monotoneOn_sub_monotoneOn` proves that a function + with locally bounded variation is the difference of two monotone functions. +* `LipschitzWith.locallyBoundedVariationOn` shows that a Lipschitz function has locally + bounded variation. + +We also give several variations around these results. + +## Implementation + +We define the variation as an extended nonnegative real, to allow for infinite variation. This makes +it possible to use the complete linear order structure of `ℝ≥0∞`. The proofs would be much +more tedious with an `ℝ`-valued or `ℝ≥0`-valued variation, since one would always need to check +that the sets one uses are nonempty and bounded above as these are only conditionally complete. +-/ + + +open scoped NNReal ENNReal Topology UniformConvergence +open Set Filter + +-- Porting note: sectioned variables because a `wlog` was broken due to extra variables in context +variable {α : Type*} [LinearOrder α] {E : Type*} [PseudoEMetricSpace E] + +/-- The (extended real valued) variation of a function `f` on a set `s` inside a linear order is +the supremum of the sum of `edist (f (u (i+1))) (f (u i))` over all finite increasing +sequences `u` in `s`. -/ +noncomputable def eVariationOn (f : α → E) (s : Set α) : ℝ≥0∞ := + ⨆ p : ℕ × { u : ℕ → α // Monotone u ∧ ∀ i, u i ∈ s }, + ∑ i ∈ Finset.range p.1, edist (f (p.2.1 (i + 1))) (f (p.2.1 i)) + +/-- A function has bounded variation on a set `s` if its total variation there is finite. -/ +def BoundedVariationOn (f : α → E) (s : Set α) := + eVariationOn f s ≠ ∞ + +/-- A function has locally bounded variation on a set `s` if, given any interval `[a, b]` with +endpoints in `s`, then the function has finite variation on `s ∩ [a, b]`. -/ +def LocallyBoundedVariationOn (f : α → E) (s : Set α) := + ∀ a b, a ∈ s → b ∈ s → BoundedVariationOn f (s ∩ Icc a b) + +/-! ## Basic computations of variation -/ + +namespace eVariationOn + +theorem nonempty_monotone_mem {s : Set α} (hs : s.Nonempty) : + Nonempty { u // Monotone u ∧ ∀ i : ℕ, u i ∈ s } := by + obtain ⟨x, hx⟩ := hs + exact ⟨⟨fun _ => x, fun i j _ => le_rfl, fun _ => hx⟩⟩ + +theorem eq_of_edist_zero_on {f f' : α → E} {s : Set α} (h : ∀ ⦃x⦄, x ∈ s → edist (f x) (f' x) = 0) : + eVariationOn f s = eVariationOn f' s := by + dsimp only [eVariationOn] + congr 1 with p : 1 + congr 1 with i : 1 + rw [edist_congr_right (h <| p.snd.prop.2 (i + 1)), edist_congr_left (h <| p.snd.prop.2 i)] + +theorem eq_of_eqOn {f f' : α → E} {s : Set α} (h : EqOn f f' s) : + eVariationOn f s = eVariationOn f' s := + eq_of_edist_zero_on fun x xs => by rw [h xs, edist_self] + +theorem sum_le (f : α → E) {s : Set α} (n : ℕ) {u : ℕ → α} (hu : Monotone u) (us : ∀ i, u i ∈ s) : + (∑ i ∈ Finset.range n, edist (f (u (i + 1))) (f (u i))) ≤ eVariationOn f s := + le_iSup_of_le ⟨n, u, hu, us⟩ le_rfl + +theorem sum_le_of_monotoneOn_Icc (f : α → E) {s : Set α} {m n : ℕ} {u : ℕ → α} + (hu : MonotoneOn u (Icc m n)) (us : ∀ i ∈ Icc m n, u i ∈ s) : + (∑ i ∈ Finset.Ico m n, edist (f (u (i + 1))) (f (u i))) ≤ eVariationOn f s := by + rcases le_total n m with hnm | hmn + · simp [Finset.Ico_eq_empty_of_le hnm] + let π := projIcc m n hmn + let v i := u (π i) + calc + ∑ i ∈ Finset.Ico m n, edist (f (u (i + 1))) (f (u i)) + = ∑ i ∈ Finset.Ico m n, edist (f (v (i + 1))) (f (v i)) := + Finset.sum_congr rfl fun i hi ↦ by + rw [Finset.mem_Ico] at hi + simp only [v, π, projIcc_of_mem hmn ⟨hi.1, hi.2.le⟩, + projIcc_of_mem hmn ⟨hi.1.trans i.le_succ, hi.2⟩] + _ ≤ ∑ i ∈ Finset.range n, edist (f (v (i + 1))) (f (v i)) := + Finset.sum_mono_set _ (Nat.Iio_eq_range ▸ Finset.Ico_subset_Iio_self) + _ ≤ eVariationOn f s := + sum_le _ _ (fun i j h ↦ hu (π i).2 (π j).2 (monotone_projIcc hmn h)) fun i ↦ us _ (π i).2 + +theorem sum_le_of_monotoneOn_Iic (f : α → E) {s : Set α} {n : ℕ} {u : ℕ → α} + (hu : MonotoneOn u (Iic n)) (us : ∀ i ≤ n, u i ∈ s) : + (∑ i ∈ Finset.range n, edist (f (u (i + 1))) (f (u i))) ≤ eVariationOn f s := by + simpa using sum_le_of_monotoneOn_Icc f (m := 0) (hu.mono Icc_subset_Iic_self) fun i hi ↦ us i hi.2 + +theorem mono (f : α → E) {s t : Set α} (hst : t ⊆ s) : eVariationOn f t ≤ eVariationOn f s := by + apply iSup_le _ + rintro ⟨n, ⟨u, hu, ut⟩⟩ + exact sum_le f n hu fun i => hst (ut i) + +theorem _root_.BoundedVariationOn.mono {f : α → E} {s : Set α} (h : BoundedVariationOn f s) + {t : Set α} (ht : t ⊆ s) : BoundedVariationOn f t := + ne_top_of_le_ne_top h (eVariationOn.mono f ht) + +theorem _root_.BoundedVariationOn.locallyBoundedVariationOn {f : α → E} {s : Set α} + (h : BoundedVariationOn f s) : LocallyBoundedVariationOn f s := fun _ _ _ _ => + h.mono inter_subset_left + +theorem edist_le (f : α → E) {s : Set α} {x y : α} (hx : x ∈ s) (hy : y ∈ s) : + edist (f x) (f y) ≤ eVariationOn f s := by + wlog hxy : y ≤ x generalizing x y + · rw [edist_comm] + exact this hy hx (le_of_not_le hxy) + let u : ℕ → α := fun n => if n = 0 then y else x + have hu : Monotone u := monotone_nat_of_le_succ fun + | 0 => hxy + | (_ + 1) => le_rfl + have us : ∀ i, u i ∈ s := fun + | 0 => hy + | (_ + 1) => hx + simpa only [Finset.sum_range_one] using sum_le f 1 hu us + +theorem eq_zero_iff (f : α → E) {s : Set α} : + eVariationOn f s = 0 ↔ ∀ x ∈ s, ∀ y ∈ s, edist (f x) (f y) = 0 := by + constructor + · rintro h x xs y ys + rw [← le_zero_iff, ← h] + exact edist_le f xs ys + · rintro h + dsimp only [eVariationOn] + rw [ENNReal.iSup_eq_zero] + rintro ⟨n, u, um, us⟩ + exact Finset.sum_eq_zero fun i _ => h _ (us i.succ) _ (us i) + +theorem constant_on {f : α → E} {s : Set α} (hf : (f '' s).Subsingleton) : + eVariationOn f s = 0 := by + rw [eq_zero_iff] + rintro x xs y ys + rw [hf ⟨x, xs, rfl⟩ ⟨y, ys, rfl⟩, edist_self] + +@[simp] +protected theorem subsingleton (f : α → E) {s : Set α} (hs : s.Subsingleton) : + eVariationOn f s = 0 := + constant_on (hs.image f) + +theorem lowerSemicontinuous_aux {ι : Type*} {F : ι → α → E} {p : Filter ι} {f : α → E} {s : Set α} + (Ffs : ∀ x ∈ s, Tendsto (fun i => F i x) p (𝓝 (f x))) {v : ℝ≥0∞} (hv : v < eVariationOn f s) : + ∀ᶠ n : ι in p, v < eVariationOn (F n) s := by + obtain ⟨⟨n, ⟨u, um, us⟩⟩, hlt⟩ : + ∃ p : ℕ × { u : ℕ → α // Monotone u ∧ ∀ i, u i ∈ s }, + v < ∑ i ∈ Finset.range p.1, edist (f ((p.2 : ℕ → α) (i + 1))) (f ((p.2 : ℕ → α) i)) := + lt_iSup_iff.mp hv + have : Tendsto (fun j => ∑ i ∈ Finset.range n, edist (F j (u (i + 1))) (F j (u i))) p + (𝓝 (∑ i ∈ Finset.range n, edist (f (u (i + 1))) (f (u i)))) := by + apply tendsto_finset_sum + exact fun i _ => Tendsto.edist (Ffs (u i.succ) (us i.succ)) (Ffs (u i) (us i)) + exact (this.eventually_const_lt hlt).mono fun i h => h.trans_le (sum_le (F i) n um us) + +/-- The map `(eVariationOn · s)` is lower semicontinuous for pointwise convergence *on `s`*. +Pointwise convergence on `s` is encoded here as uniform convergence on the family consisting of the +singletons of elements of `s`. +-/ +protected theorem lowerSemicontinuous (s : Set α) : + LowerSemicontinuous fun f : α →ᵤ[s.image singleton] E => eVariationOn f s := fun f ↦ by + apply @lowerSemicontinuous_aux _ _ _ _ (UniformOnFun α E (s.image singleton)) id (𝓝 f) f s _ + simpa only [UniformOnFun.tendsto_iff_tendstoUniformlyOn, mem_image, forall_exists_index, and_imp, + forall_apply_eq_imp_iff₂, tendstoUniformlyOn_singleton_iff_tendsto] using @tendsto_id _ (𝓝 f) + +/-- The map `(eVariationOn · s)` is lower semicontinuous for uniform convergence on `s`. -/ +theorem lowerSemicontinuous_uniformOn (s : Set α) : + LowerSemicontinuous fun f : α →ᵤ[{s}] E => eVariationOn f s := fun f ↦ by + apply @lowerSemicontinuous_aux _ _ _ _ (UniformOnFun α E {s}) id (𝓝 f) f s _ + have := @tendsto_id _ (𝓝 f) + rw [UniformOnFun.tendsto_iff_tendstoUniformlyOn] at this + simp_rw [← tendstoUniformlyOn_singleton_iff_tendsto] + exact fun x xs => (this s rfl).mono (singleton_subset_iff.mpr xs) + +theorem _root_.BoundedVariationOn.dist_le {E : Type*} [PseudoMetricSpace E] {f : α → E} + {s : Set α} (h : BoundedVariationOn f s) {x y : α} (hx : x ∈ s) (hy : y ∈ s) : + dist (f x) (f y) ≤ (eVariationOn f s).toReal := by + rw [← ENNReal.ofReal_le_ofReal_iff ENNReal.toReal_nonneg, ENNReal.ofReal_toReal h, ← edist_dist] + exact edist_le f hx hy + +theorem _root_.BoundedVariationOn.sub_le {f : α → ℝ} {s : Set α} (h : BoundedVariationOn f s) + {x y : α} (hx : x ∈ s) (hy : y ∈ s) : f x - f y ≤ (eVariationOn f s).toReal := by + apply (le_abs_self _).trans + rw [← Real.dist_eq] + exact h.dist_le hx hy + +/-- Consider a monotone function `u` parameterizing some points of a set `s`. Given `x ∈ s`, then +one can find another monotone function `v` parameterizing the same points as `u`, with `x` added. +In particular, the variation of a function along `u` is bounded by its variation along `v`. -/ +theorem add_point (f : α → E) {s : Set α} {x : α} (hx : x ∈ s) (u : ℕ → α) (hu : Monotone u) + (us : ∀ i, u i ∈ s) (n : ℕ) : + ∃ (v : ℕ → α) (m : ℕ), Monotone v ∧ (∀ i, v i ∈ s) ∧ x ∈ v '' Iio m ∧ + (∑ i ∈ Finset.range n, edist (f (u (i + 1))) (f (u i))) ≤ + ∑ j ∈ Finset.range m, edist (f (v (j + 1))) (f (v j)) := by + rcases le_or_lt (u n) x with (h | h) + · let v i := if i ≤ n then u i else x + have vs : ∀ i, v i ∈ s := fun i ↦ by + simp only [v] + split_ifs + · exact us i + · exact hx + have hv : Monotone v := by + refine monotone_nat_of_le_succ fun i => ?_ + simp only [v] + rcases lt_trichotomy i n with (hi | rfl | hi) + · have : i + 1 ≤ n := Nat.succ_le_of_lt hi + simp only [hi.le, this, if_true] + exact hu (Nat.le_succ i) + · simp only [le_refl, if_true, add_le_iff_nonpos_right, Nat.le_zero, Nat.one_ne_zero, + if_false, h] + · have A : ¬i ≤ n := hi.not_le + have B : ¬i + 1 ≤ n := fun h => A (i.le_succ.trans h) + simp only [A, B, if_false, le_rfl] + refine ⟨v, n + 2, hv, vs, (mem_image _ _ _).2 ⟨n + 1, ?_, ?_⟩, ?_⟩ + · rw [mem_Iio]; exact Nat.lt_succ_self (n + 1) + · have : ¬n + 1 ≤ n := Nat.not_succ_le_self n + simp only [v, this, ite_eq_right_iff, IsEmpty.forall_iff] + · calc + (∑ i ∈ Finset.range n, edist (f (u (i + 1))) (f (u i))) = + ∑ i ∈ Finset.range n, edist (f (v (i + 1))) (f (v i)) := by + apply Finset.sum_congr rfl fun i hi => ?_ + simp only [Finset.mem_range] at hi + have : i + 1 ≤ n := Nat.succ_le_of_lt hi + simp only [v, hi.le, this, if_true] + _ ≤ ∑ j ∈ Finset.range (n + 2), edist (f (v (j + 1))) (f (v j)) := + Finset.sum_le_sum_of_subset (Finset.range_mono (Nat.le_add_right n 2)) + have exists_N : ∃ N, N ≤ n ∧ x < u N := ⟨n, le_rfl, h⟩ + let N := Nat.find exists_N + have hN : N ≤ n ∧ x < u N := Nat.find_spec exists_N + let w : ℕ → α := fun i => if i < N then u i else if i = N then x else u (i - 1) + have ws : ∀ i, w i ∈ s := by + dsimp only [w] + intro i + split_ifs + exacts [us _, hx, us _] + have hw : Monotone w := by + apply monotone_nat_of_le_succ fun i => ?_ + dsimp only [w] + rcases lt_trichotomy (i + 1) N with (hi | hi | hi) + · have : i < N := Nat.lt_of_le_of_lt (Nat.le_succ i) hi + simp only [hi, this, if_true] + exact hu (Nat.le_succ _) + · have A : i < N := hi ▸ i.lt_succ_self + have B : ¬i + 1 < N := by rw [← hi]; exact fun h => h.ne rfl + rw [if_pos A, if_neg B, if_pos hi] + have T := Nat.find_min exists_N A + push_neg at T + exact T (A.le.trans hN.1) + · have A : ¬i < N := (Nat.lt_succ_iff.mp hi).not_lt + have B : ¬i + 1 < N := hi.not_lt + have C : ¬i + 1 = N := hi.ne.symm + have D : i + 1 - 1 = i := Nat.pred_succ i + rw [if_neg A, if_neg B, if_neg C, D] + split_ifs + · exact hN.2.le.trans (hu (le_of_not_lt A)) + · exact hu (Nat.pred_le _) + refine ⟨w, n + 1, hw, ws, (mem_image _ _ _).2 ⟨N, hN.1.trans_lt (Nat.lt_succ_self n), ?_⟩, ?_⟩ + · dsimp only [w]; rw [if_neg (lt_irrefl N), if_pos rfl] + rcases eq_or_lt_of_le (zero_le N) with (Npos | Npos) + · calc + (∑ i ∈ Finset.range n, edist (f (u (i + 1))) (f (u i))) = + ∑ i ∈ Finset.range n, edist (f (w (1 + i + 1))) (f (w (1 + i))) := by + apply Finset.sum_congr rfl fun i _hi => ?_ + dsimp only [w] + simp only [← Npos, Nat.not_lt_zero, Nat.add_succ_sub_one, add_zero, if_false, + add_eq_zero, Nat.one_ne_zero, false_and, Nat.succ_add_sub_one, zero_add] + rw [add_comm 1 i] + _ = ∑ i ∈ Finset.Ico 1 (n + 1), edist (f (w (i + 1))) (f (w i)) := by + rw [Finset.range_eq_Ico] + exact Finset.sum_Ico_add (fun i => edist (f (w (i + 1))) (f (w i))) 0 n 1 + _ ≤ ∑ j ∈ Finset.range (n + 1), edist (f (w (j + 1))) (f (w j)) := by + apply Finset.sum_le_sum_of_subset _ + rw [Finset.range_eq_Ico] + exact Finset.Ico_subset_Ico zero_le_one le_rfl + · calc + (∑ i ∈ Finset.range n, edist (f (u (i + 1))) (f (u i))) = + ((∑ i ∈ Finset.Ico 0 (N - 1), edist (f (u (i + 1))) (f (u i))) + + ∑ i ∈ Finset.Ico (N - 1) N, edist (f (u (i + 1))) (f (u i))) + + ∑ i ∈ Finset.Ico N n, edist (f (u (i + 1))) (f (u i)) := by + rw [Finset.sum_Ico_consecutive, Finset.sum_Ico_consecutive, Finset.range_eq_Ico] + · exact zero_le _ + · exact hN.1 + · exact zero_le _ + · exact Nat.pred_le _ + _ = (∑ i ∈ Finset.Ico 0 (N - 1), edist (f (w (i + 1))) (f (w i))) + + edist (f (u N)) (f (u (N - 1))) + + ∑ i ∈ Finset.Ico N n, edist (f (w (1 + i + 1))) (f (w (1 + i))) := by + congr 1 + · congr 1 + · apply Finset.sum_congr rfl fun i hi => ?_ + simp only [Finset.mem_Ico, zero_le', true_and] at hi + dsimp only [w] + have A : i + 1 < N := Nat.lt_pred_iff.1 hi + have B : i < N := Nat.lt_of_succ_lt A + rw [if_pos A, if_pos B] + · have A : N - 1 + 1 = N := Nat.succ_pred_eq_of_pos Npos + have : Finset.Ico (N - 1) N = {N - 1} := by rw [← Nat.Ico_succ_singleton, A] + simp only [this, A, Finset.sum_singleton] + · apply Finset.sum_congr rfl fun i hi => ?_ + rw [Finset.mem_Ico] at hi + dsimp only [w] + have A : ¬1 + i + 1 < N := by omega + have B : ¬1 + i + 1 = N := by omega + have C : ¬1 + i < N := by omega + have D : ¬1 + i = N := by omega + rw [if_neg A, if_neg B, if_neg C, if_neg D] + congr 3 <;> · rw [add_comm, Nat.sub_one]; apply Nat.pred_succ + _ = (∑ i ∈ Finset.Ico 0 (N - 1), edist (f (w (i + 1))) (f (w i))) + + edist (f (w (N + 1))) (f (w (N - 1))) + + ∑ i ∈ Finset.Ico (N + 1) (n + 1), edist (f (w (i + 1))) (f (w i)) := by + congr 1 + · congr 1 + · dsimp only [w] + have A : ¬N + 1 < N := Nat.not_succ_lt_self + have B : N - 1 < N := Nat.pred_lt Npos.ne' + simp only [A, not_and, not_lt, Nat.succ_ne_self, Nat.add_succ_sub_one, add_zero, + if_false, B, if_true] + · exact Finset.sum_Ico_add (fun i => edist (f (w (i + 1))) (f (w i))) N n 1 + _ ≤ ((∑ i ∈ Finset.Ico 0 (N - 1), edist (f (w (i + 1))) (f (w i))) + + ∑ i ∈ Finset.Ico (N - 1) (N + 1), edist (f (w (i + 1))) (f (w i))) + + ∑ i ∈ Finset.Ico (N + 1) (n + 1), edist (f (w (i + 1))) (f (w i)) := by + refine add_le_add (add_le_add le_rfl ?_) le_rfl + have A : N - 1 + 1 = N := Nat.succ_pred_eq_of_pos Npos + have B : N - 1 + 1 < N + 1 := A.symm ▸ N.lt_succ_self + have C : N - 1 < N + 1 := lt_of_le_of_lt N.pred_le N.lt_succ_self + rw [Finset.sum_eq_sum_Ico_succ_bot C, Finset.sum_eq_sum_Ico_succ_bot B, A, Finset.Ico_self, + Finset.sum_empty, add_zero, add_comm (edist _ _)] + exact edist_triangle _ _ _ + _ = ∑ j ∈ Finset.range (n + 1), edist (f (w (j + 1))) (f (w j)) := by + rw [Finset.sum_Ico_consecutive, Finset.sum_Ico_consecutive, Finset.range_eq_Ico] + · exact zero_le _ + · exact Nat.succ_le_succ hN.left + · exact zero_le _ + · exact N.pred_le.trans N.le_succ + +/-- The variation of a function on the union of two sets `s` and `t`, with `s` to the left of `t`, +bounds the sum of the variations along `s` and `t`. -/ +theorem add_le_union (f : α → E) {s t : Set α} (h : ∀ x ∈ s, ∀ y ∈ t, x ≤ y) : + eVariationOn f s + eVariationOn f t ≤ eVariationOn f (s ∪ t) := by + by_cases hs : s = ∅ + · simp [hs] + have : Nonempty { u // Monotone u ∧ ∀ i : ℕ, u i ∈ s } := + nonempty_monotone_mem (nonempty_iff_ne_empty.2 hs) + by_cases ht : t = ∅ + · simp [ht] + have : Nonempty { u // Monotone u ∧ ∀ i : ℕ, u i ∈ t } := + nonempty_monotone_mem (nonempty_iff_ne_empty.2 ht) + refine ENNReal.iSup_add_iSup_le ?_ + /- We start from two sequences `u` and `v` along `s` and `t` respectively, and we build a new + sequence `w` along `s ∪ t` by juxtaposing them. Its variation is larger than the sum of the + variations. -/ + rintro ⟨n, ⟨u, hu, us⟩⟩ ⟨m, ⟨v, hv, vt⟩⟩ + let w i := if i ≤ n then u i else v (i - (n + 1)) + have wst : ∀ i, w i ∈ s ∪ t := by + intro i + by_cases hi : i ≤ n + · simp [w, hi, us] + · simp [w, hi, vt] + have hw : Monotone w := by + intro i j hij + dsimp only [w] + split_ifs with h_1 h_2 h_2 + · exact hu hij + · apply h _ (us _) _ (vt _) + · exfalso; exact h_1 (hij.trans h_2) + · apply hv (tsub_le_tsub hij le_rfl) + calc + ((∑ i ∈ Finset.range n, edist (f (u (i + 1))) (f (u i))) + + ∑ i ∈ Finset.range m, edist (f (v (i + 1))) (f (v i))) = + (∑ i ∈ Finset.range n, edist (f (w (i + 1))) (f (w i))) + + ∑ i ∈ Finset.range m, edist (f (w (n + 1 + i + 1))) (f (w (n + 1 + i))) := by + dsimp only [w] + congr 1 + · refine Finset.sum_congr rfl fun i hi => ?_ + simp only [Finset.mem_range] at hi + have : i + 1 ≤ n := Nat.succ_le_of_lt hi + simp [hi.le, this] + · refine Finset.sum_congr rfl fun i hi => ?_ + simp only [Finset.mem_range] at hi + have B : ¬n + 1 + i ≤ n := by omega + have A : ¬n + 1 + i + 1 ≤ n := fun h => B ((n + 1 + i).le_succ.trans h) + have C : n + 1 + i - n = i + 1 := by + rw [tsub_eq_iff_eq_add_of_le] + · abel + · exact n.le_succ.trans (n.succ.le_add_right i) + simp only [A, B, C, Nat.succ_sub_succ_eq_sub, if_false, add_tsub_cancel_left] + _ = (∑ i ∈ Finset.range n, edist (f (w (i + 1))) (f (w i))) + + ∑ i ∈ Finset.Ico (n + 1) (n + 1 + m), edist (f (w (i + 1))) (f (w i)) := by + congr 1 + rw [Finset.range_eq_Ico] + convert Finset.sum_Ico_add (fun i : ℕ => edist (f (w (i + 1))) (f (w i))) 0 m (n + 1) + using 3 <;> abel + _ ≤ ∑ i ∈ Finset.range (n + 1 + m), edist (f (w (i + 1))) (f (w i)) := by + rw [← Finset.sum_union] + · apply Finset.sum_le_sum_of_subset _ + rintro i hi + simp only [Finset.mem_union, Finset.mem_range, Finset.mem_Ico] at hi ⊢ + cases' hi with hi hi + · exact lt_of_lt_of_le hi (n.le_succ.trans (n.succ.le_add_right m)) + · exact hi.2 + · refine Finset.disjoint_left.2 fun i hi h'i => ?_ + simp only [Finset.mem_Ico, Finset.mem_range] at hi h'i + exact hi.not_lt (Nat.lt_of_succ_le h'i.left) + _ ≤ eVariationOn f (s ∪ t) := sum_le f _ hw wst + +/-- If a set `s` is to the left of a set `t`, and both contain the boundary point `x`, then +the variation of `f` along `s ∪ t` is the sum of the variations. -/ +theorem union (f : α → E) {s t : Set α} {x : α} (hs : IsGreatest s x) (ht : IsLeast t x) : + eVariationOn f (s ∪ t) = eVariationOn f s + eVariationOn f t := by + classical + apply le_antisymm _ (eVariationOn.add_le_union f fun a ha b hb => le_trans (hs.2 ha) (ht.2 hb)) + apply iSup_le _ + rintro ⟨n, ⟨u, hu, ust⟩⟩ + obtain ⟨v, m, hv, vst, xv, huv⟩ : ∃ (v : ℕ → α) (m : ℕ), + Monotone v ∧ (∀ i, v i ∈ s ∪ t) ∧ x ∈ v '' Iio m ∧ + (∑ i ∈ Finset.range n, edist (f (u (i + 1))) (f (u i))) ≤ + ∑ j ∈ Finset.range m, edist (f (v (j + 1))) (f (v j)) := + eVariationOn.add_point f (mem_union_left t hs.1) u hu ust n + obtain ⟨N, hN, Nx⟩ : ∃ N, N < m ∧ v N = x := xv + calc + (∑ j ∈ Finset.range n, edist (f (u (j + 1))) (f (u j))) ≤ + ∑ j ∈ Finset.range m, edist (f (v (j + 1))) (f (v j)) := + huv + _ = (∑ j ∈ Finset.Ico 0 N, edist (f (v (j + 1))) (f (v j))) + + ∑ j ∈ Finset.Ico N m, edist (f (v (j + 1))) (f (v j)) := by + rw [Finset.range_eq_Ico, Finset.sum_Ico_consecutive _ (zero_le _) hN.le] + _ ≤ eVariationOn f s + eVariationOn f t := by + refine add_le_add ?_ ?_ + · apply sum_le_of_monotoneOn_Icc _ (hv.monotoneOn _) fun i hi => ?_ + rcases vst i with (h | h); · exact h + have : v i = x := by + apply le_antisymm + · rw [← Nx]; exact hv hi.2 + · exact ht.2 h + rw [this] + exact hs.1 + · apply sum_le_of_monotoneOn_Icc _ (hv.monotoneOn _) fun i hi => ?_ + rcases vst i with (h | h); swap; · exact h + have : v i = x := by + apply le_antisymm + · exact hs.2 h + · rw [← Nx]; exact hv hi.1 + rw [this] + exact ht.1 + +theorem Icc_add_Icc (f : α → E) {s : Set α} {a b c : α} (hab : a ≤ b) (hbc : b ≤ c) (hb : b ∈ s) : + eVariationOn f (s ∩ Icc a b) + eVariationOn f (s ∩ Icc b c) = eVariationOn f (s ∩ Icc a c) := by + have A : IsGreatest (s ∩ Icc a b) b := + ⟨⟨hb, hab, le_rfl⟩, inter_subset_right.trans Icc_subset_Iic_self⟩ + have B : IsLeast (s ∩ Icc b c) b := + ⟨⟨hb, le_rfl, hbc⟩, inter_subset_right.trans Icc_subset_Ici_self⟩ + rw [← eVariationOn.union f A B, ← inter_union_distrib_left, Icc_union_Icc_eq_Icc hab hbc] + +section Monotone + +variable {β : Type*} [LinearOrder β] + +theorem comp_le_of_monotoneOn (f : α → E) {s : Set α} {t : Set β} (φ : β → α) (hφ : MonotoneOn φ t) + (φst : MapsTo φ t s) : eVariationOn (f ∘ φ) t ≤ eVariationOn f s := + iSup_le fun ⟨n, u, hu, ut⟩ => + le_iSup_of_le ⟨n, φ ∘ u, fun x y xy => hφ (ut x) (ut y) (hu xy), fun i => φst (ut i)⟩ le_rfl + +theorem comp_le_of_antitoneOn (f : α → E) {s : Set α} {t : Set β} (φ : β → α) (hφ : AntitoneOn φ t) + (φst : MapsTo φ t s) : eVariationOn (f ∘ φ) t ≤ eVariationOn f s := by + refine iSup_le ?_ + rintro ⟨n, u, hu, ut⟩ + rw [← Finset.sum_range_reflect] + refine (Finset.sum_congr rfl fun x hx => ?_).trans_le <| le_iSup_of_le + ⟨n, fun i => φ (u <| n - i), fun x y xy => hφ (ut _) (ut _) (hu <| Nat.sub_le_sub_left xy n), + fun i => φst (ut _)⟩ + le_rfl + rw [Finset.mem_range] at hx + dsimp only [Subtype.coe_mk, Function.comp_apply] + rw [edist_comm] + congr 4 <;> omega + +theorem comp_eq_of_monotoneOn (f : α → E) {t : Set β} (φ : β → α) (hφ : MonotoneOn φ t) : + eVariationOn (f ∘ φ) t = eVariationOn f (φ '' t) := by + apply le_antisymm (comp_le_of_monotoneOn f φ hφ (mapsTo_image φ t)) + cases isEmpty_or_nonempty β + · convert zero_le (_ : ℝ≥0∞) + exact eVariationOn.subsingleton f <| + (subsingleton_of_subsingleton.image _).anti (surjOn_image φ t) + let ψ := φ.invFunOn t + have ψφs : EqOn (φ ∘ ψ) id (φ '' t) := (surjOn_image φ t).rightInvOn_invFunOn + have ψts : MapsTo ψ (φ '' t) t := (surjOn_image φ t).mapsTo_invFunOn + have hψ : MonotoneOn ψ (φ '' t) := Function.monotoneOn_of_rightInvOn_of_mapsTo hφ ψφs ψts + change eVariationOn (f ∘ id) (φ '' t) ≤ eVariationOn (f ∘ φ) t + rw [← eq_of_eqOn (ψφs.comp_left : EqOn (f ∘ φ ∘ ψ) (f ∘ id) (φ '' t))] + exact comp_le_of_monotoneOn _ ψ hψ ψts + +theorem comp_inter_Icc_eq_of_monotoneOn (f : α → E) {t : Set β} (φ : β → α) (hφ : MonotoneOn φ t) + {x y : β} (hx : x ∈ t) (hy : y ∈ t) : + eVariationOn (f ∘ φ) (t ∩ Icc x y) = eVariationOn f (φ '' t ∩ Icc (φ x) (φ y)) := by + rcases le_total x y with (h | h) + · convert comp_eq_of_monotoneOn f φ (hφ.mono Set.inter_subset_left) + apply le_antisymm + · rintro _ ⟨⟨u, us, rfl⟩, vφx, vφy⟩ + rcases le_total x u with (xu | ux) + · rcases le_total u y with (uy | yu) + · exact ⟨u, ⟨us, ⟨xu, uy⟩⟩, rfl⟩ + · rw [le_antisymm vφy (hφ hy us yu)] + exact ⟨y, ⟨hy, ⟨h, le_rfl⟩⟩, rfl⟩ + · rw [← le_antisymm vφx (hφ us hx ux)] + exact ⟨x, ⟨hx, ⟨le_rfl, h⟩⟩, rfl⟩ + · rintro _ ⟨u, ⟨⟨hu, xu, uy⟩, rfl⟩⟩ + exact ⟨⟨u, hu, rfl⟩, ⟨hφ hx hu xu, hφ hu hy uy⟩⟩ + · rw [eVariationOn.subsingleton, eVariationOn.subsingleton] + exacts [(Set.subsingleton_Icc_of_ge (hφ hy hx h)).anti Set.inter_subset_right, + (Set.subsingleton_Icc_of_ge h).anti Set.inter_subset_right] + +theorem comp_eq_of_antitoneOn (f : α → E) {t : Set β} (φ : β → α) (hφ : AntitoneOn φ t) : + eVariationOn (f ∘ φ) t = eVariationOn f (φ '' t) := by + apply le_antisymm (comp_le_of_antitoneOn f φ hφ (mapsTo_image φ t)) + cases isEmpty_or_nonempty β + · convert zero_le (_ : ℝ≥0∞) + exact eVariationOn.subsingleton f <| (subsingleton_of_subsingleton.image _).anti + (surjOn_image φ t) + let ψ := φ.invFunOn t + have ψφs : EqOn (φ ∘ ψ) id (φ '' t) := (surjOn_image φ t).rightInvOn_invFunOn + have ψts := (surjOn_image φ t).mapsTo_invFunOn + have hψ : AntitoneOn ψ (φ '' t) := Function.antitoneOn_of_rightInvOn_of_mapsTo hφ ψφs ψts + change eVariationOn (f ∘ id) (φ '' t) ≤ eVariationOn (f ∘ φ) t + rw [← eq_of_eqOn (ψφs.comp_left : EqOn (f ∘ φ ∘ ψ) (f ∘ id) (φ '' t))] + exact comp_le_of_antitoneOn _ ψ hψ ψts + +open OrderDual + +theorem comp_ofDual (f : α → E) (s : Set α) : + eVariationOn (f ∘ ofDual) (ofDual ⁻¹' s) = eVariationOn f s := by + convert comp_eq_of_antitoneOn f ofDual fun _ _ _ _ => id + simp only [Equiv.image_preimage] + +end Monotone + +end eVariationOn + +/-! ## Monotone functions and bounded variation -/ + +theorem MonotoneOn.eVariationOn_le {f : α → ℝ} {s : Set α} (hf : MonotoneOn f s) {a b : α} + (as : a ∈ s) (bs : b ∈ s) : eVariationOn f (s ∩ Icc a b) ≤ ENNReal.ofReal (f b - f a) := by + apply iSup_le _ + rintro ⟨n, ⟨u, hu, us⟩⟩ + calc + (∑ i ∈ Finset.range n, edist (f (u (i + 1))) (f (u i))) = + ∑ i ∈ Finset.range n, ENNReal.ofReal (f (u (i + 1)) - f (u i)) := by + refine Finset.sum_congr rfl fun i hi => ?_ + simp only [Finset.mem_range] at hi + rw [edist_dist, Real.dist_eq, abs_of_nonneg] + exact sub_nonneg_of_le (hf (us i).1 (us (i + 1)).1 (hu (Nat.le_succ _))) + _ = ENNReal.ofReal (∑ i ∈ Finset.range n, (f (u (i + 1)) - f (u i))) := by + rw [ENNReal.ofReal_sum_of_nonneg] + intro i _ + exact sub_nonneg_of_le (hf (us i).1 (us (i + 1)).1 (hu (Nat.le_succ _))) + _ = ENNReal.ofReal (f (u n) - f (u 0)) := by rw [Finset.sum_range_sub fun i => f (u i)] + _ ≤ ENNReal.ofReal (f b - f a) := by + apply ENNReal.ofReal_le_ofReal + exact sub_le_sub (hf (us n).1 bs (us n).2.2) (hf as (us 0).1 (us 0).2.1) + +theorem MonotoneOn.locallyBoundedVariationOn {f : α → ℝ} {s : Set α} (hf : MonotoneOn f s) : + LocallyBoundedVariationOn f s := fun _ _ as bs => + ((hf.eVariationOn_le as bs).trans_lt ENNReal.ofReal_lt_top).ne + +/-- The **signed** variation of `f` on the interval `Icc a b` intersected with the set `s`, +squashed to a real (therefore only really meaningful if the variation is finite) +-/ +noncomputable def variationOnFromTo (f : α → E) (s : Set α) (a b : α) : ℝ := + if a ≤ b then (eVariationOn f (s ∩ Icc a b)).toReal else -(eVariationOn f (s ∩ Icc b a)).toReal + +namespace variationOnFromTo + +variable (f : α → E) (s : Set α) + +protected theorem self (a : α) : variationOnFromTo f s a a = 0 := by + dsimp only [variationOnFromTo] + rw [if_pos le_rfl, Icc_self, eVariationOn.subsingleton, ENNReal.zero_toReal] + exact fun x hx y hy => hx.2.trans hy.2.symm + +protected theorem nonneg_of_le {a b : α} (h : a ≤ b) : 0 ≤ variationOnFromTo f s a b := by + simp only [variationOnFromTo, if_pos h, ENNReal.toReal_nonneg] + +protected theorem eq_neg_swap (a b : α) : + variationOnFromTo f s a b = -variationOnFromTo f s b a := by + rcases lt_trichotomy a b with (ab | rfl | ba) + · simp only [variationOnFromTo, if_pos ab.le, if_neg ab.not_le, neg_neg] + · simp only [variationOnFromTo.self, neg_zero] + · simp only [variationOnFromTo, if_pos ba.le, if_neg ba.not_le, neg_neg] + +protected theorem nonpos_of_ge {a b : α} (h : b ≤ a) : variationOnFromTo f s a b ≤ 0 := by + rw [variationOnFromTo.eq_neg_swap] + exact neg_nonpos_of_nonneg (variationOnFromTo.nonneg_of_le f s h) + +protected theorem eq_of_le {a b : α} (h : a ≤ b) : + variationOnFromTo f s a b = (eVariationOn f (s ∩ Icc a b)).toReal := + if_pos h + +protected theorem eq_of_ge {a b : α} (h : b ≤ a) : + variationOnFromTo f s a b = -(eVariationOn f (s ∩ Icc b a)).toReal := by + rw [variationOnFromTo.eq_neg_swap, neg_inj, variationOnFromTo.eq_of_le f s h] + +protected theorem add {f : α → E} {s : Set α} (hf : LocallyBoundedVariationOn f s) {a b c : α} + (ha : a ∈ s) (hb : b ∈ s) (hc : c ∈ s) : + variationOnFromTo f s a b + variationOnFromTo f s b c = variationOnFromTo f s a c := by + symm + refine additive_of_isTotal ((· : α) ≤ ·) (variationOnFromTo f s) (· ∈ s) ?_ ?_ ha hb hc + · rintro x y _xs _ys + simp only [variationOnFromTo.eq_neg_swap f s y x, Subtype.coe_mk, add_neg_cancel, + forall_true_left] + · rintro x y z xy yz xs ys zs + rw [variationOnFromTo.eq_of_le f s xy, variationOnFromTo.eq_of_le f s yz, + variationOnFromTo.eq_of_le f s (xy.trans yz), + ← ENNReal.toReal_add (hf x y xs ys) (hf y z ys zs), eVariationOn.Icc_add_Icc f xy yz ys] + +variable {f s} in +protected theorem edist_zero_of_eq_zero (hf : LocallyBoundedVariationOn f s) + {a b : α} (ha : a ∈ s) (hb : b ∈ s) (h : variationOnFromTo f s a b = 0) : + edist (f a) (f b) = 0 := by + wlog h' : a ≤ b + · rw [edist_comm] + apply this hf hb ha _ (le_of_not_le h') + rw [variationOnFromTo.eq_neg_swap, h, neg_zero] + · apply le_antisymm _ (zero_le _) + rw [← ENNReal.ofReal_zero, ← h, variationOnFromTo.eq_of_le f s h', + ENNReal.ofReal_toReal (hf a b ha hb)] + apply eVariationOn.edist_le + exacts [⟨ha, ⟨le_rfl, h'⟩⟩, ⟨hb, ⟨h', le_rfl⟩⟩] + +protected theorem eq_left_iff {f : α → E} {s : Set α} (hf : LocallyBoundedVariationOn f s) + {a b c : α} (ha : a ∈ s) (hb : b ∈ s) (hc : c ∈ s) : + variationOnFromTo f s a b = variationOnFromTo f s a c ↔ variationOnFromTo f s b c = 0 := by + simp only [← variationOnFromTo.add hf ha hb hc, self_eq_add_right] + +protected theorem eq_zero_iff_of_le {f : α → E} {s : Set α} (hf : LocallyBoundedVariationOn f s) + {a b : α} (ha : a ∈ s) (hb : b ∈ s) (ab : a ≤ b) : + variationOnFromTo f s a b = 0 ↔ + ∀ ⦃x⦄ (_hx : x ∈ s ∩ Icc a b) ⦃y⦄ (_hy : y ∈ s ∩ Icc a b), edist (f x) (f y) = 0 := by + rw [variationOnFromTo.eq_of_le _ _ ab, ENNReal.toReal_eq_zero_iff, or_iff_left (hf a b ha hb), + eVariationOn.eq_zero_iff] + +protected theorem eq_zero_iff_of_ge {f : α → E} {s : Set α} (hf : LocallyBoundedVariationOn f s) + {a b : α} (ha : a ∈ s) (hb : b ∈ s) (ba : b ≤ a) : + variationOnFromTo f s a b = 0 ↔ + ∀ ⦃x⦄ (_hx : x ∈ s ∩ Icc b a) ⦃y⦄ (_hy : y ∈ s ∩ Icc b a), edist (f x) (f y) = 0 := by + rw [variationOnFromTo.eq_of_ge _ _ ba, neg_eq_zero, ENNReal.toReal_eq_zero_iff, + or_iff_left (hf b a hb ha), eVariationOn.eq_zero_iff] + +protected theorem eq_zero_iff {f : α → E} {s : Set α} (hf : LocallyBoundedVariationOn f s) {a b : α} + (ha : a ∈ s) (hb : b ∈ s) : + variationOnFromTo f s a b = 0 ↔ + ∀ ⦃x⦄ (_hx : x ∈ s ∩ uIcc a b) ⦃y⦄ (_hy : y ∈ s ∩ uIcc a b), edist (f x) (f y) = 0 := by + rcases le_total a b with (ab | ba) + · rw [uIcc_of_le ab] + exact variationOnFromTo.eq_zero_iff_of_le hf ha hb ab + · rw [uIcc_of_ge ba] + exact variationOnFromTo.eq_zero_iff_of_ge hf ha hb ba + +variable {f} {s} + +protected theorem monotoneOn (hf : LocallyBoundedVariationOn f s) {a : α} (as : a ∈ s) : + MonotoneOn (variationOnFromTo f s a) s := by + rintro b bs c cs bc + rw [← variationOnFromTo.add hf as bs cs] + exact le_add_of_nonneg_right (variationOnFromTo.nonneg_of_le f s bc) + +protected theorem antitoneOn (hf : LocallyBoundedVariationOn f s) {b : α} (bs : b ∈ s) : + AntitoneOn (fun a => variationOnFromTo f s a b) s := by + rintro a as c cs ac + dsimp only + rw [← variationOnFromTo.add hf as cs bs] + exact le_add_of_nonneg_left (variationOnFromTo.nonneg_of_le f s ac) + +protected theorem sub_self_monotoneOn {f : α → ℝ} {s : Set α} (hf : LocallyBoundedVariationOn f s) + {a : α} (as : a ∈ s) : MonotoneOn (variationOnFromTo f s a - f) s := by + rintro b bs c cs bc + rw [Pi.sub_apply, Pi.sub_apply, le_sub_iff_add_le, add_comm_sub, ← le_sub_iff_add_le'] + calc + f c - f b ≤ |f c - f b| := le_abs_self _ + _ = dist (f b) (f c) := by rw [dist_comm, Real.dist_eq] + _ ≤ variationOnFromTo f s b c := by + rw [variationOnFromTo.eq_of_le f s bc, dist_edist] + apply ENNReal.toReal_mono (hf b c bs cs) + apply eVariationOn.edist_le f + exacts [⟨bs, le_rfl, bc⟩, ⟨cs, bc, le_rfl⟩] + _ = variationOnFromTo f s a c - variationOnFromTo f s a b := by + rw [← variationOnFromTo.add hf as bs cs, add_sub_cancel_left] + +protected theorem comp_eq_of_monotoneOn {β : Type*} [LinearOrder β] (f : α → E) {t : Set β} + (φ : β → α) (hφ : MonotoneOn φ t) {x y : β} (hx : x ∈ t) (hy : y ∈ t) : + variationOnFromTo (f ∘ φ) t x y = variationOnFromTo f (φ '' t) (φ x) (φ y) := by + rcases le_total x y with (h | h) + · rw [variationOnFromTo.eq_of_le _ _ h, variationOnFromTo.eq_of_le _ _ (hφ hx hy h), + eVariationOn.comp_inter_Icc_eq_of_monotoneOn f φ hφ hx hy] + · rw [variationOnFromTo.eq_of_ge _ _ h, variationOnFromTo.eq_of_ge _ _ (hφ hy hx h), + eVariationOn.comp_inter_Icc_eq_of_monotoneOn f φ hφ hy hx] + +end variationOnFromTo + +/-- If a real valued function has bounded variation on a set, then it is a difference of monotone +functions there. -/ +theorem LocallyBoundedVariationOn.exists_monotoneOn_sub_monotoneOn {f : α → ℝ} {s : Set α} + (h : LocallyBoundedVariationOn f s) : + ∃ p q : α → ℝ, MonotoneOn p s ∧ MonotoneOn q s ∧ f = p - q := by + rcases eq_empty_or_nonempty s with (rfl | ⟨c, cs⟩) + · exact ⟨f, 0, subsingleton_empty.monotoneOn _, subsingleton_empty.monotoneOn _, + (sub_zero f).symm⟩ + · exact ⟨_, _, variationOnFromTo.monotoneOn h cs, variationOnFromTo.sub_self_monotoneOn h cs, + (sub_sub_cancel _ _).symm⟩ + +/-! ## Lipschitz functions and bounded variation -/ + +section LipschitzOnWith + +variable {F : Type*} [PseudoEMetricSpace F] + +theorem LipschitzOnWith.comp_eVariationOn_le {f : E → F} {C : ℝ≥0} {t : Set E} + (h : LipschitzOnWith C f t) {g : α → E} {s : Set α} (hg : MapsTo g s t) : + eVariationOn (f ∘ g) s ≤ C * eVariationOn g s := by + apply iSup_le _ + rintro ⟨n, ⟨u, hu, us⟩⟩ + calc + (∑ i ∈ Finset.range n, edist (f (g (u (i + 1)))) (f (g (u i)))) ≤ + ∑ i ∈ Finset.range n, C * edist (g (u (i + 1))) (g (u i)) := + Finset.sum_le_sum fun i _ => h (hg (us _)) (hg (us _)) + _ = C * ∑ i ∈ Finset.range n, edist (g (u (i + 1))) (g (u i)) := by rw [Finset.mul_sum] + _ ≤ C * eVariationOn g s := mul_le_mul_left' (eVariationOn.sum_le _ _ hu us) _ + +theorem LipschitzOnWith.comp_boundedVariationOn {f : E → F} {C : ℝ≥0} {t : Set E} + (hf : LipschitzOnWith C f t) {g : α → E} {s : Set α} (hg : MapsTo g s t) + (h : BoundedVariationOn g s) : BoundedVariationOn (f ∘ g) s := + ne_top_of_le_ne_top (ENNReal.mul_ne_top ENNReal.coe_ne_top h) (hf.comp_eVariationOn_le hg) + +theorem LipschitzOnWith.comp_locallyBoundedVariationOn {f : E → F} {C : ℝ≥0} {t : Set E} + (hf : LipschitzOnWith C f t) {g : α → E} {s : Set α} (hg : MapsTo g s t) + (h : LocallyBoundedVariationOn g s) : LocallyBoundedVariationOn (f ∘ g) s := + fun x y xs ys => + hf.comp_boundedVariationOn (hg.mono_left inter_subset_left) (h x y xs ys) + +theorem LipschitzWith.comp_boundedVariationOn {f : E → F} {C : ℝ≥0} (hf : LipschitzWith C f) + {g : α → E} {s : Set α} (h : BoundedVariationOn g s) : BoundedVariationOn (f ∘ g) s := + hf.lipschitzOnWith.comp_boundedVariationOn (mapsTo_univ _ _) h + +theorem LipschitzWith.comp_locallyBoundedVariationOn {f : E → F} {C : ℝ≥0} + (hf : LipschitzWith C f) {g : α → E} {s : Set α} (h : LocallyBoundedVariationOn g s) : + LocallyBoundedVariationOn (f ∘ g) s := + hf.lipschitzOnWith.comp_locallyBoundedVariationOn (mapsTo_univ _ _) h + +theorem LipschitzOnWith.locallyBoundedVariationOn {f : ℝ → E} {C : ℝ≥0} {s : Set ℝ} + (hf : LipschitzOnWith C f s) : LocallyBoundedVariationOn f s := + hf.comp_locallyBoundedVariationOn (mapsTo_id _) + (@monotoneOn_id ℝ _ s).locallyBoundedVariationOn + +theorem LipschitzWith.locallyBoundedVariationOn {f : ℝ → E} {C : ℝ≥0} (hf : LipschitzWith C f) + (s : Set ℝ) : LocallyBoundedVariationOn f s := + hf.lipschitzOnWith.locallyBoundedVariationOn + +end LipschitzOnWith From 0e6b027f3e22586bd249beb2326be105277eded6 Mon Sep 17 00:00:00 2001 From: peabrainiac Date: Mon, 9 Dec 2024 10:07:27 +0000 Subject: [PATCH 589/829] feat(Geometry/Manifold): manifolds are locally path-connected (#17142) Shows that spaces modelled on locally path-connected spaces are themselves locally path-connected, and that the common model spaces `EuclideanHalfSpace n` and `EuclideanQuadrant n` are locally path-connected. Co-authored-by: peabrainiac <43812953+peabrainiac@users.noreply.github.com> --- .../Complex/UpperHalfPlane/Topology.lean | 2 +- Mathlib/Analysis/Convex/Normed.lean | 4 ---- Mathlib/Geometry/Manifold/ChartedSpace.lean | 18 ++++++++++++++++ Mathlib/Geometry/Manifold/Instances/Real.lean | 21 +++++++++++++++++++ .../Algebra/Module/LocallyConvex.lean | 19 +++++++++++++---- .../Compactness/DeltaGeneratedSpace.lean | 2 +- Mathlib/Topology/Connected/PathConnected.lean | 5 +++++ 7 files changed, 61 insertions(+), 10 deletions(-) diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean index ac416004b29f9..2925cad492d23 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean @@ -5,7 +5,7 @@ Authors: Yury Kudryashov -/ import Mathlib.Analysis.Complex.UpperHalfPlane.Basic import Mathlib.Analysis.Convex.Contractible -import Mathlib.Analysis.Convex.Normed +import Mathlib.Analysis.LocallyConvex.WithSeminorms import Mathlib.Analysis.Complex.Convex import Mathlib.Analysis.Complex.ReImTopology import Mathlib.Topology.Homotopy.Contractible diff --git a/Mathlib/Analysis/Convex/Normed.lean b/Mathlib/Analysis/Convex/Normed.lean index 9a305c1457de0..2277ac898fcf3 100644 --- a/Mathlib/Analysis/Convex/Normed.lean +++ b/Mathlib/Analysis/Convex/Normed.lean @@ -111,10 +111,6 @@ theorem isBounded_convexHull {s : Set E} : instance (priority := 100) NormedSpace.instPathConnectedSpace : PathConnectedSpace E := TopologicalAddGroup.pathConnectedSpace -instance (priority := 100) NormedSpace.instLocPathConnectedSpace : LocPathConnectedSpace E := - .of_bases (fun _ => Metric.nhds_basis_ball) fun x r r_pos => - (convex_ball x r).isPathConnected <| by simp [r_pos] - theorem Wbtw.dist_add_dist {x y z : P} (h : Wbtw ℝ x y z) : dist x y + dist y z = dist x z := by obtain ⟨a, ⟨ha₀, ha₁⟩, rfl⟩ := h diff --git a/Mathlib/Geometry/Manifold/ChartedSpace.lean b/Mathlib/Geometry/Manifold/ChartedSpace.lean index 2c3d9ef063c40..0e455da5da7e2 100644 --- a/Mathlib/Geometry/Manifold/ChartedSpace.lean +++ b/Mathlib/Geometry/Manifold/ChartedSpace.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import Mathlib.Topology.PartialHomeomorph +import Mathlib.Topology.Connected.PathConnected /-! # Charted spaces @@ -693,6 +694,23 @@ theorem ChartedSpace.locallyConnectedSpace [LocallyConnectedSpace H] : LocallyCo · rintro x s ⟨⟨-, -, hsconn⟩, hssubset⟩ exact hsconn.isPreconnected.image _ ((e x).continuousOn_symm.mono hssubset) +/-- If a topological space `M` admits an atlas with locally path-connected charts, + then `M` itself is locally path-connected. -/ +theorem ChartedSpace.locPathConnectedSpace [LocPathConnectedSpace H] : LocPathConnectedSpace M := by + refine ⟨fun x ↦ ⟨fun s ↦ ⟨fun hs ↦ ?_, fun ⟨u, hu⟩ ↦ Filter.mem_of_superset hu.1.1 hu.2⟩⟩⟩ + let e := chartAt H x + let t := s ∩ e.source + have ht : t ∈ 𝓝 x := Filter.inter_mem hs (chart_source_mem_nhds _ _) + refine ⟨e.symm '' pathComponentIn (e x) (e '' t), ⟨?_, ?_⟩, (?_ : _ ⊆ t).trans inter_subset_left⟩ + · nth_rewrite 1 [← e.left_inv (mem_chart_source _ _)] + apply e.symm.image_mem_nhds (by simp [e]) + exact pathComponentIn_mem_nhds <| e.image_mem_nhds (mem_chart_source _ _) ht + · refine (isPathConnected_pathComponentIn <| mem_image_of_mem e (mem_of_mem_nhds ht)).image' ?_ + refine e.continuousOn_symm.mono <| subset_trans ?_ e.map_source'' + exact (pathComponentIn_mono <| image_mono inter_subset_right).trans pathComponentIn_subset + · exact (image_mono pathComponentIn_subset).trans + (PartialEquiv.symm_image_image_of_subset_source _ inter_subset_right).subset + /-- If `M` is modelled on `H'` and `H'` is itself modelled on `H`, then we can consider `M` as being modelled on `H`. -/ def ChartedSpace.comp (H : Type*) [TopologicalSpace H] (H' : Type*) [TopologicalSpace H'] diff --git a/Mathlib/Geometry/Manifold/Instances/Real.lean b/Mathlib/Geometry/Manifold/Instances/Real.lean index 7f10c097d74fe..299a57f94ff35 100644 --- a/Mathlib/Geometry/Manifold/Instances/Real.lean +++ b/Mathlib/Geometry/Manifold/Instances/Real.lean @@ -87,6 +87,27 @@ theorem EuclideanHalfSpace.ext [NeZero n] (x y : EuclideanHalfSpace n) (h : x.1 = y.1) : x = y := Subtype.eq h +theorem EuclideanHalfSpace.convex [NeZero n] : + Convex ℝ { x : EuclideanSpace ℝ (Fin n) | 0 ≤ x 0 } := + fun _ hx _ hy _ _ _ _ _ ↦ by dsimp at hx hy ⊢; positivity + +theorem EuclideanQuadrant.convex : + Convex ℝ { x : EuclideanSpace ℝ (Fin n) | ∀ i, 0 ≤ x i } := + fun _ hx _ hy _ _ _ _ _ i ↦ by dsimp at hx hy ⊢; specialize hx i; specialize hy i; positivity + +instance EuclideanHalfSpace.pathConnectedSpace [NeZero n] : + PathConnectedSpace (EuclideanHalfSpace n) := + isPathConnected_iff_pathConnectedSpace.mp <| convex.isPathConnected ⟨0, by simp⟩ + +instance EuclideanQuadrant.pathConnectedSpace : PathConnectedSpace (EuclideanQuadrant n) := + isPathConnected_iff_pathConnectedSpace.mp <| convex.isPathConnected ⟨0, by simp⟩ + +instance [NeZero n] : LocPathConnectedSpace (EuclideanHalfSpace n) := + EuclideanHalfSpace.convex.locPathConnectedSpace + +instance : LocPathConnectedSpace (EuclideanQuadrant n) := + EuclideanQuadrant.convex.locPathConnectedSpace + theorem range_euclideanHalfSpace (n : ℕ) [NeZero n] : (range fun x : EuclideanHalfSpace n => x.val) = { y | 0 ≤ y 0 } := Subtype.range_val diff --git a/Mathlib/Topology/Algebra/Module/LocallyConvex.lean b/Mathlib/Topology/Algebra/Module/LocallyConvex.lean index dbed63cbce723..8d72fea405dff 100644 --- a/Mathlib/Topology/Algebra/Module/LocallyConvex.lean +++ b/Mathlib/Topology/Algebra/Module/LocallyConvex.lean @@ -91,10 +91,21 @@ theorem locallyConvexSpace_iff_exists_convex_subset_zero : (locallyConvexSpace_iff_zero 𝕜 E).trans hasBasis_self -- see Note [lower instance priority] -instance (priority := 100) LocallyConvexSpace.toLocallyConnectedSpace [Module ℝ E] - [ContinuousSMul ℝ E] [LocallyConvexSpace ℝ E] : LocallyConnectedSpace E := - locallyConnectedSpace_of_connected_bases _ _ - (fun x => @LocallyConvexSpace.convex_basis ℝ _ _ _ _ _ _ x) fun _ _ hs => hs.2.isPreconnected +instance (priority := 100) LocallyConvexSpace.toLocPathConnectedSpace [Module ℝ E] + [ContinuousSMul ℝ E] [LocallyConvexSpace ℝ E] : LocPathConnectedSpace E := + .of_bases (fun x ↦ convex_basis (𝕜 := ℝ) x) + fun _ _ hs ↦ hs.2.isPathConnected <| nonempty_of_mem <| mem_of_mem_nhds hs.1 + +/-- Convex subsets of locally convex spaces are locally path-connected. -/ +theorem Convex.locPathConnectedSpace [Module ℝ E] [ContinuousSMul ℝ E] [LocallyConvexSpace ℝ E] + {S : Set E} (hS : Convex ℝ S) : LocPathConnectedSpace S := by + refine ⟨fun x ↦ ⟨fun s ↦ ⟨fun hs ↦ ?_, fun ⟨t, ht⟩ ↦ mem_of_superset ht.1.1 ht.2⟩⟩⟩ + let ⟨t, ht⟩ := (mem_nhds_subtype S x s).mp hs + let ⟨t', ht'⟩ := (LocallyConvexSpace.convex_basis (𝕜 := ℝ) x.1).mem_iff.mp ht.1 + refine ⟨(↑) ⁻¹' t', ⟨?_, ?_⟩, (preimage_mono ht'.2).trans ht.2⟩ + · exact continuousAt_subtype_val.preimage_mem_nhds ht'.1.1 + · refine Subtype.preimage_coe_self_inter _ _ ▸ IsPathConnected.preimage_coe ?_ inter_subset_left + exact (hS.inter ht'.1.2).isPathConnected ⟨x, x.2, mem_of_mem_nhds ht'.1.1⟩ end Module diff --git a/Mathlib/Topology/Compactness/DeltaGeneratedSpace.lean b/Mathlib/Topology/Compactness/DeltaGeneratedSpace.lean index c7c9e34c6a6ae..889946c028e5c 100644 --- a/Mathlib/Topology/Compactness/DeltaGeneratedSpace.lean +++ b/Mathlib/Topology/Compactness/DeltaGeneratedSpace.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Ben Eltschig. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Ben Eltschig -/ -import Mathlib.Analysis.Convex.Normed +import Mathlib.Analysis.LocallyConvex.WithSeminorms /-! # Delta-generated topological spaces diff --git a/Mathlib/Topology/Connected/PathConnected.lean b/Mathlib/Topology/Connected/PathConnected.lean index c5d94601f96f7..6f82df2380b21 100644 --- a/Mathlib/Topology/Connected/PathConnected.lean +++ b/Mathlib/Topology/Connected/PathConnected.lean @@ -1185,6 +1185,11 @@ protected theorem IsClosed.pathComponent (x : X) : IsClosed (pathComponent x) := protected theorem IsClopen.pathComponent (x : X) : IsClopen (pathComponent x) := ⟨.pathComponent x, .pathComponent x⟩ +lemma pathComponentIn_mem_nhds (hF : F ∈ 𝓝 x) : pathComponentIn x F ∈ 𝓝 x := by + let ⟨u, huF, hu, hxu⟩ := mem_nhds_iff.mp hF + exact mem_nhds_iff.mpr ⟨pathComponentIn x u, pathComponentIn_mono huF, + hu.pathComponentIn x, mem_pathComponentIn_self hxu⟩ + theorem pathConnectedSpace_iff_connectedSpace : PathConnectedSpace X ↔ ConnectedSpace X := by refine ⟨fun _ ↦ inferInstance, fun h ↦ ⟨inferInstance, fun x y ↦ ?_⟩⟩ rw [← mem_pathComponent_iff, (IsClopen.pathComponent _).eq_univ] <;> simp From a388e3f2b2f88a7a961915598f7f142c4c673e6c Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 9 Dec 2024 10:26:22 +0000 Subject: [PATCH 590/829] chore: delete Mathlib.Util.IncludeStr (#19828) There's no need to deprecate; the only effect of this code is to cause an ambiguity error with the Lean builtin version. --- Mathlib.lean | 1 - Mathlib/Util/IncludeStr.lean | 22 ---------------------- 2 files changed, 23 deletions(-) delete mode 100644 Mathlib/Util/IncludeStr.lean diff --git a/Mathlib.lean b/Mathlib.lean index e287ba642ed7e..e58d390587fbf 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -5400,7 +5400,6 @@ import Mathlib.Util.DischargerAsTactic import Mathlib.Util.Export import Mathlib.Util.FormatTable import Mathlib.Util.GetAllModules -import Mathlib.Util.IncludeStr import Mathlib.Util.LongNames import Mathlib.Util.MemoFix import Mathlib.Util.Notation3 diff --git a/Mathlib/Util/IncludeStr.lean b/Mathlib/Util/IncludeStr.lean deleted file mode 100644 index 662a23839c5ff..0000000000000 --- a/Mathlib/Util/IncludeStr.lean +++ /dev/null @@ -1,22 +0,0 @@ -/- -Copyright (c) 2021 Henrik Böving. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Henrik Böving, Xubai Wang --/ -import Mathlib.Init - -/-! -# Defines the `include_str` macro. --/ - -namespace Mathlib.Util - -/-- A term macro that includes the content of a file, as a string. -/ -elab (name := includeStr) "include_str " str:str : term => do - let some str := str.1.isStrLit? | Lean.Elab.throwUnsupportedSyntax - let srcPath := System.FilePath.mk (← Lean.MonadLog.getFileName) - let some srcDir := srcPath.parent | throwError "{srcPath} not in a valid directory" - let path := srcDir / str - Lean.mkStrLit <$> IO.FS.readFile path - -end Mathlib.Util From 888b1efecaf4a3dbf1ba69ffdb2934452650a19f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 9 Dec 2024 11:16:00 +0000 Subject: [PATCH 591/829] refactor: unbundle `Multiset.card` (#19777) This is a first step towards not importing algebra to define multisets. --- Mathlib/Algebra/BigOperators/Finsupp.lean | 3 +++ .../Algebra/BigOperators/Group/Finset.lean | 4 ++++ .../Combinatorics/Additive/FreimanHom.lean | 8 ++++---- .../Combinatorics/Optimization/ValuedCSP.lean | 3 +-- Mathlib/Data/Multiset/Basic.lean | 20 +++++++++++-------- Mathlib/Data/Sym/Basic.lean | 4 ++-- .../NormalizedFactors.lean | 2 +- 7 files changed, 27 insertions(+), 17 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Finsupp.lean b/Mathlib/Algebra/BigOperators/Finsupp.lean index dd91da8f00899..b992f81eade0b 100644 --- a/Mathlib/Algebra/BigOperators/Finsupp.lean +++ b/Mathlib/Algebra/BigOperators/Finsupp.lean @@ -573,6 +573,9 @@ theorem Finsupp.mul_sum (b : S) (s : α →₀ R) {f : α → R → S} : end +@[simp] lemma Multiset.card_finsuppSum [Zero M] (f : ι →₀ M) (g : ι → M → Multiset α) : + card (f.sum g) = f.sum fun i m ↦ card (g i m) := map_finsupp_sum cardHom .. + namespace Nat -- Porting note: Needed to replace pow with (· ^ ·) diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index 34cd0b147e819..e6b48d65deaa9 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -2089,6 +2089,10 @@ end List namespace Multiset +@[simp] +lemma card_sum (s : Finset ι) (f : ι → Multiset α) : card (∑ i ∈ s, f i) = ∑ i ∈ s, card (f i) := + map_sum cardHom .. + theorem disjoint_list_sum_left {a : Multiset α} {l : List (Multiset α)} : Disjoint l.sum a ↔ ∀ b ∈ l, Disjoint b a := by induction l with diff --git a/Mathlib/Combinatorics/Additive/FreimanHom.lean b/Mathlib/Combinatorics/Additive/FreimanHom.lean index 804e026e94c5a..b3ec72a21f901 100644 --- a/Mathlib/Combinatorics/Additive/FreimanHom.lean +++ b/Mathlib/Combinatorics/Additive/FreimanHom.lean @@ -284,8 +284,8 @@ lemma IsMulFreimanHom.mono (hmn : m ≤ n) (hf : IsMulFreimanHom n A B f) : obtain ha | ha := ha · exact htA ha · rwa [eq_of_mem_replicate ha] - · rw [_root_.map_add, card_replicate, hs, Nat.add_sub_cancel' hmn] - · rw [_root_.map_add, card_replicate, ht, Nat.add_sub_cancel' hmn] + · rw [card_add, card_replicate, hs, Nat.add_sub_cancel' hmn] + · rw [card_add, card_replicate, ht, Nat.add_sub_cancel' hmn] · rw [prod_add, prod_add, h] end CancelCommMonoid @@ -317,8 +317,8 @@ lemma IsMulFreimanIso.mono {hmn : m ≤ n} (hf : IsMulFreimanIso n A B f) : obtain ha | ha := ha · exact htA ha · rwa [eq_of_mem_replicate ha] - · rw [_root_.map_add, card_replicate, hs, Nat.add_sub_cancel' hmn] - · rw [_root_.map_add, card_replicate, ht, Nat.add_sub_cancel' hmn] + · rw [card_add, card_replicate, hs, Nat.add_sub_cancel' hmn] + · rw [card_add, card_replicate, ht, Nat.add_sub_cancel' hmn] end CancelCommMonoid diff --git a/Mathlib/Combinatorics/Optimization/ValuedCSP.lean b/Mathlib/Combinatorics/Optimization/ValuedCSP.lean index 956feba606f04..25a39b29065a3 100644 --- a/Mathlib/Combinatorics/Optimization/ValuedCSP.lean +++ b/Mathlib/Combinatorics/Optimization/ValuedCSP.lean @@ -90,8 +90,7 @@ variable {m : ℕ} /-- Arity of the "output" of the fractional operation. -/ @[simp] -def FractionalOperation.size (ω : FractionalOperation D m) : ℕ := - Multiset.card.toFun ω +def FractionalOperation.size (ω : FractionalOperation D m) : ℕ := ω.card /-- Fractional operation is valid iff nonempty. -/ def FractionalOperation.IsValid (ω : FractionalOperation D m) : Prop := diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index 83aedd9344065..473e920a37352 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -665,10 +665,7 @@ theorem nsmul_cons {s : Multiset α} (n : ℕ) (a : α) : /-- The cardinality of a multiset is the sum of the multiplicities of all its elements, or simply the length of the underlying list. -/ -def card : Multiset α →+ ℕ where - toFun s := (Quot.liftOn s length) fun _l₁ _l₂ => Perm.length_eq - map_zero' := rfl - map_add' s t := Quotient.inductionOn₂ s t length_append +def card (s : Multiset α) : ℕ := Quot.liftOn s length fun _l₁ _l₂ => Perm.length_eq @[simp] theorem coe_card (l : List α) : card (l : Multiset α) = length l := @@ -682,11 +679,18 @@ theorem length_toList (s : Multiset α) : s.toList.length = card s := by theorem card_zero : @card α 0 = 0 := rfl -theorem card_add (s t : Multiset α) : card (s + t) = card s + card t := - card.map_add s t +@[simp] lemma card_add (s t : Multiset α) : card (s + t) = card s + card t := + Quotient.inductionOn₂ s t length_append -theorem card_nsmul (s : Multiset α) (n : ℕ) : card (n • s) = n * card s := by - rw [card.map_nsmul s n, Nat.nsmul_eq_mul] +/-- `Multiset.card` bundled as a monoid hom. -/ +@[simps] +def cardHom : Multiset α →+ ℕ where + toFun := card + map_zero' := card_zero + map_add' := card_add + +@[simp] +lemma card_nsmul (s : Multiset α) (n : ℕ) : card (n • s) = n * card s := cardHom.map_nsmul .. @[simp] theorem card_cons (a : α) (s : Multiset α) : card (a ::ₘ s) = card s + 1 := diff --git a/Mathlib/Data/Sym/Basic.lean b/Mathlib/Data/Sym/Basic.lean index 32aca18c1b5c4..4daed5e6c93fa 100644 --- a/Mathlib/Data/Sym/Basic.lean +++ b/Mathlib/Data/Sym/Basic.lean @@ -385,7 +385,7 @@ def equivCongr (e : α ≃ β) : Sym α n ≃ Sym β n where /-- "Attach" a proof that `a ∈ s` to each element `a` in `s` to produce an element of the symmetric power on `{x // x ∈ s}`. -/ def attach (s : Sym α n) : Sym { x // x ∈ s } n := - ⟨s.val.attach, by (conv_rhs => rw [← s.2, ← Multiset.card_attach]); rfl⟩ + ⟨s.val.attach, by (conv_rhs => rw [← s.2, ← Multiset.card_attach])⟩ @[simp] theorem attach_mk {m : Multiset α} {hc : Multiset.card m = n} : @@ -441,7 +441,7 @@ theorem mem_cast (h : n = m) : a ∈ Sym.cast h s ↔ a ∈ s := /-- Append a pair of `Sym` terms. -/ def append (s : Sym α n) (s' : Sym α n') : Sym α (n + n') := - ⟨s.1 + s'.1, by rw [map_add, s.2, s'.2]⟩ + ⟨s.1 + s'.1, by rw [Multiset.card_add, s.2, s'.2]⟩ @[simp] theorem append_inj_right (s : Sym α n) {t t' : Sym α n'} : s.append t = s.append t' ↔ t = t' := diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/NormalizedFactors.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/NormalizedFactors.lean index 3d30df0c3b813..4db1bb549e738 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain/NormalizedFactors.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/NormalizedFactors.lean @@ -222,7 +222,7 @@ theorem mem_normalizedFactors_iff [Subsingleton αˣ] {p x : α} (hx : x ≠ 0) theorem exists_associated_prime_pow_of_unique_normalized_factor {p r : α} (h : ∀ {m}, m ∈ normalizedFactors r → m = p) (hr : r ≠ 0) : ∃ i : ℕ, Associated (p ^ i) r := by - use Multiset.card.toFun (normalizedFactors r) + use (normalizedFactors r).card have := UniqueFactorizationMonoid.normalizedFactors_prod hr rwa [Multiset.eq_replicate_of_mem fun b => h, Multiset.prod_replicate] at this From dda658e6181761ecabf97f47d001ffd08287d1ea Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 9 Dec 2024 11:53:04 +0000 Subject: [PATCH 592/829] =?UTF-8?q?feat(List):=20`count=20a=20(l=E2=82=81.?= =?UTF-8?q?diff=20l=E2=82=82)=20=3D=20count=20a=20l=E2=82=81=20-=20count?= =?UTF-8?q?=20a=20l=E2=82=82`=20(#19778)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This will help not importing algebra when defining multisets --- Mathlib/Data/List/Count.lean | 26 +++++++++++++++++++------- Mathlib/Data/List/Perm/Subperm.lean | 5 +++++ 2 files changed, 24 insertions(+), 7 deletions(-) diff --git a/Mathlib/Data/List/Count.lean b/Mathlib/Data/List/Count.lean index 13c127b563b2a..c9608f6a4c96f 100644 --- a/Mathlib/Data/List/Count.lean +++ b/Mathlib/Data/List/Count.lean @@ -13,9 +13,8 @@ This file proves basic properties of `List.countP` and `List.count`, which count elements of a list satisfying a predicate and equal to a given element respectively. -/ +assert_not_exists Monoid assert_not_exists Set.range -assert_not_exists GroupWithZero -assert_not_exists Ring open Nat @@ -23,9 +22,24 @@ variable {α : Type*} namespace List -/-! ### count -/ - -section Count +lemma countP_erase [DecidableEq α] (p : α → Bool) (l : List α) (a : α) : + countP p (l.erase a) = countP p l - if a ∈ l ∧ p a then 1 else 0 := by + rw [countP_eq_length_filter, countP_eq_length_filter, ← erase_filter, length_erase] + aesop + +lemma count_diff [DecidableEq α] (a : α) (l₁ : List α) : + ∀ l₂, count a (l₁.diff l₂) = count a l₁ - count a l₂ + | [] => rfl + | b :: l₂ => by + simp only [diff_cons, count_diff, count_erase, beq_iff_eq, Nat.sub_right_comm, count_cons, + Nat.sub_add_eq] + +lemma countP_diff [DecidableEq α] {l₁ l₂ : List α} (hl : l₂ <+~ l₁) (p : α → Bool) : + countP p (l₁.diff l₂) = countP p l₁ - countP p l₂ := by + refine (Nat.sub_eq_of_eq_add ?_).symm + rw [← countP_append] + exact ((subperm_append_diff_self_of_count_le <| subperm_ext_iff.1 hl).symm.trans + perm_append_comm).countP_eq _ @[simp] theorem count_map_of_injective {β} [DecidableEq α] [DecidableEq β] (l : List α) (f : α → β) @@ -34,6 +48,4 @@ theorem count_map_of_injective {β} [DecidableEq α] [DecidableEq β] (l : List unfold Function.comp simp only [hf.beq_eq] -end Count - end List diff --git a/Mathlib/Data/List/Perm/Subperm.lean b/Mathlib/Data/List/Perm/Subperm.lean index bc574cf55a74d..3266017ace4f2 100644 --- a/Mathlib/Data/List/Perm/Subperm.lean +++ b/Mathlib/Data/List/Perm/Subperm.lean @@ -30,6 +30,11 @@ attribute [trans] Subperm.trans end Subperm +/-- See also `List.subperm_ext_iff`. -/ +lemma subperm_iff_count [DecidableEq α] : l₁ <+~ l₂ ↔ ∀ a, count a l₁ ≤ count a l₂ := + subperm_ext_iff.trans <| forall_congr' fun a ↦ by + by_cases ha : a ∈ l₁ <;> simp [ha, count_eq_zero_of_not_mem] + lemma subperm_iff : l₁ <+~ l₂ ↔ ∃ l, l ~ l₂ ∧ l₁ <+ l := by refine ⟨?_, fun ⟨l, h₁, h₂⟩ ↦ h₂.subperm.trans h₁.subperm⟩ rintro ⟨l, h₁, h₂⟩ From fc349f0c3fe125fa6018ac1f30f93dc4b706b460 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 9 Dec 2024 12:57:46 +0000 Subject: [PATCH 593/829] chore: don't need group-theoretic exponent to set up finite dimensional vector spaces (#19827) --- Mathlib.lean | 4 +- Mathlib/Algebra/Module/FreeLocus.lean | 1 + .../Polynomial/Module/FiniteDimensional.lean | 1 + .../Analysis/InnerProductSpace/Rayleigh.lean | 1 + Mathlib/FieldTheory/Adjoin.lean | 1 + Mathlib/FieldTheory/Finite/Polynomial.lean | 1 + .../SplittingField/Construction.lean | 1 + .../LinearAlgebra/Dimension/DivisionRing.lean | 121 +-------------- .../Dimension/ErdosKaplansky.lean | 144 ++++++++++++++++++ Mathlib/LinearAlgebra/Dimension/Finite.lean | 14 -- .../{Torsion.lean => Torsion/Basic.lean} | 0 .../Dimension/Torsion/Finite.lean | 39 +++++ Mathlib/LinearAlgebra/Dual.lean | 3 +- Mathlib/LinearAlgebra/Eigenspace/Basic.lean | 3 + Mathlib/LinearAlgebra/FiniteDimensional.lean | 4 + .../FreeModule/Finite/Matrix.lean | 1 + Mathlib/LinearAlgebra/FreeModule/Norm.lean | 1 + Mathlib/LinearAlgebra/Matrix/Diagonal.lean | 1 + .../Matrix/NonsingularInverse.lean | 1 + Mathlib/LinearAlgebra/Reflection.lean | 1 + Mathlib/LinearAlgebra/Semisimple.lean | 1 + .../Covering/BesicovitchVectorSpace.lean | 1 + .../Integral/IntervalIntegral.lean | 1 + .../NumberTheory/NumberField/Units/Basic.lean | 1 + Mathlib/RepresentationTheory/Basic.lean | 1 + Mathlib/RingTheory/Artinian.lean | 2 + Mathlib/RingTheory/DedekindDomain/Ideal.lean | 1 + Mathlib/RingTheory/LocalRing/Module.lean | 4 +- Mathlib/RingTheory/Polynomial/GaussLemma.lean | 1 - Mathlib/RingTheory/PowerBasis.lean | 1 + 30 files changed, 219 insertions(+), 138 deletions(-) create mode 100644 Mathlib/LinearAlgebra/Dimension/ErdosKaplansky.lean rename Mathlib/LinearAlgebra/Dimension/{Torsion.lean => Torsion/Basic.lean} (100%) create mode 100644 Mathlib/LinearAlgebra/Dimension/Torsion/Finite.lean diff --git a/Mathlib.lean b/Mathlib.lean index e58d390587fbf..f6c3ebb239a3b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3316,6 +3316,7 @@ import Mathlib.LinearAlgebra.Determinant import Mathlib.LinearAlgebra.Dimension.Basic import Mathlib.LinearAlgebra.Dimension.Constructions import Mathlib.LinearAlgebra.Dimension.DivisionRing +import Mathlib.LinearAlgebra.Dimension.ErdosKaplansky import Mathlib.LinearAlgebra.Dimension.Finite import Mathlib.LinearAlgebra.Dimension.Finrank import Mathlib.LinearAlgebra.Dimension.Free @@ -3324,7 +3325,8 @@ import Mathlib.LinearAlgebra.Dimension.LinearMap import Mathlib.LinearAlgebra.Dimension.Localization import Mathlib.LinearAlgebra.Dimension.RankNullity import Mathlib.LinearAlgebra.Dimension.StrongRankCondition -import Mathlib.LinearAlgebra.Dimension.Torsion +import Mathlib.LinearAlgebra.Dimension.Torsion.Basic +import Mathlib.LinearAlgebra.Dimension.Torsion.Finite import Mathlib.LinearAlgebra.DirectSum.Finsupp import Mathlib.LinearAlgebra.DirectSum.TensorProduct import Mathlib.LinearAlgebra.Dual diff --git a/Mathlib/Algebra/Module/FreeLocus.lean b/Mathlib/Algebra/Module/FreeLocus.lean index 1487895fd48a4..68a8d2d46a780 100644 --- a/Mathlib/Algebra/Module/FreeLocus.lean +++ b/Mathlib/Algebra/Module/FreeLocus.lean @@ -10,6 +10,7 @@ import Mathlib.RingTheory.LocalRing.Module import Mathlib.RingTheory.Localization.Free import Mathlib.RingTheory.Localization.LocalizationLocalization import Mathlib.Topology.LocallyConstant.Basic +import Mathlib.RingTheory.TensorProduct.Free /-! diff --git a/Mathlib/Algebra/Polynomial/Module/FiniteDimensional.lean b/Mathlib/Algebra/Polynomial/Module/FiniteDimensional.lean index faca138037327..0918dde79efba 100644 --- a/Mathlib/Algebra/Polynomial/Module/FiniteDimensional.lean +++ b/Mathlib/Algebra/Polynomial/Module/FiniteDimensional.lean @@ -5,6 +5,7 @@ Authors: Oliver Nash -/ import Mathlib.FieldTheory.Minpoly.Field import Mathlib.Algebra.Polynomial.Module.AEval +import Mathlib.Algebra.Module.Torsion /-! # Polynomial modules in finite dimensions diff --git a/Mathlib/Analysis/InnerProductSpace/Rayleigh.lean b/Mathlib/Analysis/InnerProductSpace/Rayleigh.lean index a3620202f0233..ce57eea606617 100644 --- a/Mathlib/Analysis/InnerProductSpace/Rayleigh.lean +++ b/Mathlib/Analysis/InnerProductSpace/Rayleigh.lean @@ -8,6 +8,7 @@ import Mathlib.Analysis.InnerProductSpace.Dual import Mathlib.Analysis.InnerProductSpace.Adjoint import Mathlib.Analysis.Calculus.LagrangeMultipliers import Mathlib.LinearAlgebra.Eigenspace.Basic +import Mathlib.Algebra.EuclideanDomain.Basic /-! # The Rayleigh quotient diff --git a/Mathlib/FieldTheory/Adjoin.lean b/Mathlib/FieldTheory/Adjoin.lean index 51d8f9c940bb7..54c839c356865 100644 --- a/Mathlib/FieldTheory/Adjoin.lean +++ b/Mathlib/FieldTheory/Adjoin.lean @@ -12,6 +12,7 @@ import Mathlib.RingTheory.Adjoin.Dimension import Mathlib.RingTheory.Finiteness.TensorProduct import Mathlib.RingTheory.TensorProduct.Basic import Mathlib.SetTheory.Cardinal.Subfield +import Mathlib.LinearAlgebra.Dual /-! # Adjoining Elements to Fields diff --git a/Mathlib/FieldTheory/Finite/Polynomial.lean b/Mathlib/FieldTheory/Finite/Polynomial.lean index fa34749883456..d93f2b2225060 100644 --- a/Mathlib/FieldTheory/Finite/Polynomial.lean +++ b/Mathlib/FieldTheory/Finite/Polynomial.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.MvPolynomial.Expand import Mathlib.FieldTheory.Finite.Basic import Mathlib.LinearAlgebra.FiniteDimensional import Mathlib.RingTheory.MvPolynomial.Basic +import Mathlib.LinearAlgebra.Dual /-! ## Polynomials over finite fields diff --git a/Mathlib/FieldTheory/SplittingField/Construction.lean b/Mathlib/FieldTheory/SplittingField/Construction.lean index 9177d7e31ecfa..5d6be060b9747 100644 --- a/Mathlib/FieldTheory/SplittingField/Construction.lean +++ b/Mathlib/FieldTheory/SplittingField/Construction.lean @@ -6,6 +6,7 @@ Authors: Chris Hughes import Mathlib.Algebra.CharP.Algebra import Mathlib.FieldTheory.SplittingField.IsSplittingField import Mathlib.RingTheory.Algebraic.Basic +import Mathlib.LinearAlgebra.Dual /-! # Splitting fields diff --git a/Mathlib/LinearAlgebra/Dimension/DivisionRing.lean b/Mathlib/LinearAlgebra/Dimension/DivisionRing.lean index 775bb233d9c48..19bb2b8c796ae 100644 --- a/Mathlib/LinearAlgebra/Dimension/DivisionRing.lean +++ b/Mathlib/LinearAlgebra/Dimension/DivisionRing.lean @@ -6,7 +6,6 @@ Kim Morrison, Chris Hughes, Anne Baanen, Junyan Xu -/ import Mathlib.LinearAlgebra.Basis.VectorSpace import Mathlib.LinearAlgebra.Dimension.Finite -import Mathlib.SetTheory.Cardinal.Subfield import Mathlib.LinearAlgebra.Dimension.RankNullity /-! @@ -22,9 +21,8 @@ For vector spaces (i.e. modules over a field), we have * `rank_quotient_add_rank_of_divisionRing`: if `V₁` is a submodule of `V`, then `Module.rank (V/V₁) + Module.rank V₁ = Module.rank V`. * `rank_range_add_rank_ker`: the rank-nullity theorem. -* `rank_dual_eq_card_dual_of_aleph0_le_rank`: The **Erdős-Kaplansky Theorem** which says that - the dimension of an infinite-dimensional dual space over a division ring has dimension - equal to its cardinality. + +See also `Mathlib.LinearAlgebra.Dimension.ErdosKaplansky` for the Erdős-Kaplansky theorem. -/ @@ -217,118 +215,3 @@ noncomputable def setBasisOfTopLeSpanOfCardEqFinrank {s : Set V} [Fintype s] (_root_.trans s.toFinset_card.symm card_eq) end Basis - -section Cardinal - -variable (K) -variable [DivisionRing K] - -/-- Key lemma towards the Erdős-Kaplansky theorem from https://mathoverflow.net/a/168624 -/ -theorem max_aleph0_card_le_rank_fun_nat : max ℵ₀ #K ≤ Module.rank K (ℕ → K) := by - have aleph0_le : ℵ₀ ≤ Module.rank K (ℕ → K) := (rank_finsupp_self K ℕ).symm.trans_le - (Finsupp.lcoeFun.rank_le_of_injective <| by exact DFunLike.coe_injective) - refine max_le aleph0_le ?_ - obtain card_K | card_K := le_or_lt #K ℵ₀ - · exact card_K.trans aleph0_le - by_contra! - obtain ⟨⟨ιK, bK⟩⟩ := Module.Free.exists_basis (R := K) (M := ℕ → K) - let L := Subfield.closure (Set.range (fun i : ιK × ℕ ↦ bK i.1 i.2)) - have hLK : #L < #K := by - refine (Subfield.cardinalMk_closure_le_max _).trans_lt - (max_lt_iff.mpr ⟨mk_range_le.trans_lt ?_, card_K⟩) - rwa [mk_prod, ← aleph0, lift_uzero, bK.mk_eq_rank'', mul_aleph0_eq aleph0_le] - letI := Module.compHom K (RingHom.op L.subtype) - obtain ⟨⟨ιL, bL⟩⟩ := Module.Free.exists_basis (R := Lᵐᵒᵖ) (M := K) - have card_ιL : ℵ₀ ≤ #ιL := by - contrapose! hLK - haveI := @Fintype.ofFinite _ (lt_aleph0_iff_finite.mp hLK) - rw [bL.repr.toEquiv.cardinal_eq, mk_finsupp_of_fintype, - ← MulOpposite.opEquiv.cardinal_eq] at card_K ⊢ - apply power_nat_le - contrapose! card_K - exact (power_lt_aleph0 card_K <| nat_lt_aleph0 _).le - obtain ⟨e⟩ := lift_mk_le'.mp (card_ιL.trans_eq (lift_uzero #ιL).symm) - have rep_e := bK.linearCombination_repr (bL ∘ e) - rw [Finsupp.linearCombination_apply, Finsupp.sum] at rep_e - set c := bK.repr (bL ∘ e) - set s := c.support - let f i (j : s) : L := ⟨bK j i, Subfield.subset_closure ⟨(j, i), rfl⟩⟩ - have : ¬LinearIndependent Lᵐᵒᵖ f := fun h ↦ by - have := h.cardinal_lift_le_rank - rw [lift_uzero, (LinearEquiv.piCongrRight fun _ ↦ MulOpposite.opLinearEquiv Lᵐᵒᵖ).rank_eq, - rank_fun'] at this - exact (nat_lt_aleph0 _).not_le this - obtain ⟨t, g, eq0, i, hi, hgi⟩ := not_linearIndependent_iff.mp this - refine hgi (linearIndependent_iff'.mp (bL.linearIndependent.comp e e.injective) t g ?_ i hi) - clear_value c s - simp_rw [← rep_e, Finset.sum_apply, Pi.smul_apply, Finset.smul_sum] - rw [Finset.sum_comm] - refine Finset.sum_eq_zero fun i hi ↦ ?_ - replace eq0 := congr_arg L.subtype (congr_fun eq0 ⟨i, hi⟩) - rw [Finset.sum_apply, map_sum] at eq0 - have : SMulCommClass Lᵐᵒᵖ K K := ⟨fun _ _ _ ↦ mul_assoc _ _ _⟩ - simp_rw [smul_comm _ (c i), ← Finset.smul_sum] - erw [eq0, smul_zero] - -variable {K} - -open Function in -theorem rank_fun_infinite {ι : Type v} [hι : Infinite ι] : Module.rank K (ι → K) = #(ι → K) := by - obtain ⟨⟨ιK, bK⟩⟩ := Module.Free.exists_basis (R := K) (M := ι → K) - obtain ⟨e⟩ := lift_mk_le'.mp ((aleph0_le_mk_iff.mpr hι).trans_eq (lift_uzero #ι).symm) - have := LinearMap.lift_rank_le_of_injective _ <| - LinearMap.funLeft_injective_of_surjective K K _ (invFun_surjective e.injective) - rw [lift_umax.{u,v}, lift_id'.{u,v}] at this - have key := (lift_le.{v}.mpr <| max_aleph0_card_le_rank_fun_nat K).trans this - rw [lift_max, lift_aleph0, max_le_iff] at key - haveI : Infinite ιK := by - rw [← aleph0_le_mk_iff, bK.mk_eq_rank'']; exact key.1 - rw [bK.repr.toEquiv.cardinal_eq, mk_finsupp_lift_of_infinite, - lift_umax.{u,v}, lift_id'.{u,v}, bK.mk_eq_rank'', eq_comm, max_eq_left] - exact key.2 - -/-- The **Erdős-Kaplansky Theorem**: the dual of an infinite-dimensional vector space - over a division ring has dimension equal to its cardinality. -/ -theorem rank_dual_eq_card_dual_of_aleph0_le_rank' {V : Type*} [AddCommGroup V] [Module K V] - (h : ℵ₀ ≤ Module.rank K V) : Module.rank Kᵐᵒᵖ (V →ₗ[K] K) = #(V →ₗ[K] K) := by - obtain ⟨⟨ι, b⟩⟩ := Module.Free.exists_basis (R := K) (M := V) - rw [← b.mk_eq_rank'', aleph0_le_mk_iff] at h - have e := (b.constr Kᵐᵒᵖ (M' := K)).symm.trans - (LinearEquiv.piCongrRight fun _ ↦ MulOpposite.opLinearEquiv Kᵐᵒᵖ) - rw [e.rank_eq, e.toEquiv.cardinal_eq] - apply rank_fun_infinite - -/-- The **Erdős-Kaplansky Theorem** over a field. -/ -theorem rank_dual_eq_card_dual_of_aleph0_le_rank {K V} [Field K] [AddCommGroup V] [Module K V] - (h : ℵ₀ ≤ Module.rank K V) : Module.rank K (V →ₗ[K] K) = #(V →ₗ[K] K) := by - obtain ⟨⟨ι, b⟩⟩ := Module.Free.exists_basis (R := K) (M := V) - rw [← b.mk_eq_rank'', aleph0_le_mk_iff] at h - have e := (b.constr K (M' := K)).symm - rw [e.rank_eq, e.toEquiv.cardinal_eq] - apply rank_fun_infinite - -theorem lift_rank_lt_rank_dual' {V : Type v} [AddCommGroup V] [Module K V] - (h : ℵ₀ ≤ Module.rank K V) : - Cardinal.lift.{u} (Module.rank K V) < Module.rank Kᵐᵒᵖ (V →ₗ[K] K) := by - obtain ⟨⟨ι, b⟩⟩ := Module.Free.exists_basis (R := K) (M := V) - rw [← b.mk_eq_rank'', rank_dual_eq_card_dual_of_aleph0_le_rank' h, - ← (b.constr ℕ (M' := K)).toEquiv.cardinal_eq, mk_arrow] - apply cantor' - erw [nat_lt_lift_iff, one_lt_iff_nontrivial] - infer_instance - -theorem lift_rank_lt_rank_dual {K : Type u} {V : Type v} [Field K] [AddCommGroup V] [Module K V] - (h : ℵ₀ ≤ Module.rank K V) : - Cardinal.lift.{u} (Module.rank K V) < Module.rank K (V →ₗ[K] K) := by - rw [rank_dual_eq_card_dual_of_aleph0_le_rank h, ← rank_dual_eq_card_dual_of_aleph0_le_rank' h] - exact lift_rank_lt_rank_dual' h - -theorem rank_lt_rank_dual' {V : Type u} [AddCommGroup V] [Module K V] (h : ℵ₀ ≤ Module.rank K V) : - Module.rank K V < Module.rank Kᵐᵒᵖ (V →ₗ[K] K) := by - convert lift_rank_lt_rank_dual' h; rw [lift_id] - -theorem rank_lt_rank_dual {K V : Type u} [Field K] [AddCommGroup V] [Module K V] - (h : ℵ₀ ≤ Module.rank K V) : Module.rank K V < Module.rank K (V →ₗ[K] K) := by - convert lift_rank_lt_rank_dual h; rw [lift_id] - -end Cardinal diff --git a/Mathlib/LinearAlgebra/Dimension/ErdosKaplansky.lean b/Mathlib/LinearAlgebra/Dimension/ErdosKaplansky.lean new file mode 100644 index 0000000000000..a8f0c42dc1fbe --- /dev/null +++ b/Mathlib/LinearAlgebra/Dimension/ErdosKaplansky.lean @@ -0,0 +1,144 @@ +/- +Copyright (c) 2018 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro, Johannes Hölzl, Sander Dahmen, +Kim Morrison, Chris Hughes, Anne Baanen, Junyan Xu +-/ +import Mathlib.Algebra.Field.Opposite +import Mathlib.LinearAlgebra.Basis.VectorSpace +import Mathlib.LinearAlgebra.Dimension.Constructions +import Mathlib.SetTheory.Cardinal.Subfield + +/-! +# Erdős-Kaplansky theorem + +For modules over a division ring, we have + +* `rank_dual_eq_card_dual_of_aleph0_le_rank`: The **Erdős-Kaplansky Theorem** which says that + the dimension of an infinite-dimensional dual space over a division ring has dimension + equal to its cardinality. + +-/ + +noncomputable section + +universe u v + +variable {K : Type u} + +open Cardinal + +section Cardinal + +variable (K) +variable [DivisionRing K] + +/-- Key lemma towards the Erdős-Kaplansky theorem from https://mathoverflow.net/a/168624 -/ +theorem max_aleph0_card_le_rank_fun_nat : max ℵ₀ #K ≤ Module.rank K (ℕ → K) := by + have aleph0_le : ℵ₀ ≤ Module.rank K (ℕ → K) := (rank_finsupp_self K ℕ).symm.trans_le + (Finsupp.lcoeFun.rank_le_of_injective <| by exact DFunLike.coe_injective) + refine max_le aleph0_le ?_ + obtain card_K | card_K := le_or_lt #K ℵ₀ + · exact card_K.trans aleph0_le + by_contra! + obtain ⟨⟨ιK, bK⟩⟩ := Module.Free.exists_basis (R := K) (M := ℕ → K) + let L := Subfield.closure (Set.range (fun i : ιK × ℕ ↦ bK i.1 i.2)) + have hLK : #L < #K := by + refine (Subfield.cardinalMk_closure_le_max _).trans_lt + (max_lt_iff.mpr ⟨mk_range_le.trans_lt ?_, card_K⟩) + rwa [mk_prod, ← aleph0, lift_uzero, bK.mk_eq_rank'', mul_aleph0_eq aleph0_le] + letI := Module.compHom K (RingHom.op L.subtype) + obtain ⟨⟨ιL, bL⟩⟩ := Module.Free.exists_basis (R := Lᵐᵒᵖ) (M := K) + have card_ιL : ℵ₀ ≤ #ιL := by + contrapose! hLK + haveI := @Fintype.ofFinite _ (lt_aleph0_iff_finite.mp hLK) + rw [bL.repr.toEquiv.cardinal_eq, mk_finsupp_of_fintype, + ← MulOpposite.opEquiv.cardinal_eq] at card_K ⊢ + apply power_nat_le + contrapose! card_K + exact (power_lt_aleph0 card_K <| nat_lt_aleph0 _).le + obtain ⟨e⟩ := lift_mk_le'.mp (card_ιL.trans_eq (lift_uzero #ιL).symm) + have rep_e := bK.linearCombination_repr (bL ∘ e) + rw [Finsupp.linearCombination_apply, Finsupp.sum] at rep_e + set c := bK.repr (bL ∘ e) + set s := c.support + let f i (j : s) : L := ⟨bK j i, Subfield.subset_closure ⟨(j, i), rfl⟩⟩ + have : ¬LinearIndependent Lᵐᵒᵖ f := fun h ↦ by + have := h.cardinal_lift_le_rank + rw [lift_uzero, (LinearEquiv.piCongrRight fun _ ↦ MulOpposite.opLinearEquiv Lᵐᵒᵖ).rank_eq, + rank_fun'] at this + exact (nat_lt_aleph0 _).not_le this + obtain ⟨t, g, eq0, i, hi, hgi⟩ := not_linearIndependent_iff.mp this + refine hgi (linearIndependent_iff'.mp (bL.linearIndependent.comp e e.injective) t g ?_ i hi) + clear_value c s + simp_rw [← rep_e, Finset.sum_apply, Pi.smul_apply, Finset.smul_sum] + rw [Finset.sum_comm] + refine Finset.sum_eq_zero fun i hi ↦ ?_ + replace eq0 := congr_arg L.subtype (congr_fun eq0 ⟨i, hi⟩) + rw [Finset.sum_apply, map_sum] at eq0 + have : SMulCommClass Lᵐᵒᵖ K K := ⟨fun _ _ _ ↦ mul_assoc _ _ _⟩ + simp_rw [smul_comm _ (c i), ← Finset.smul_sum] + erw [eq0, smul_zero] + +variable {K} + +open Function in +theorem rank_fun_infinite {ι : Type v} [hι : Infinite ι] : Module.rank K (ι → K) = #(ι → K) := by + obtain ⟨⟨ιK, bK⟩⟩ := Module.Free.exists_basis (R := K) (M := ι → K) + obtain ⟨e⟩ := lift_mk_le'.mp ((aleph0_le_mk_iff.mpr hι).trans_eq (lift_uzero #ι).symm) + have := LinearMap.lift_rank_le_of_injective _ <| + LinearMap.funLeft_injective_of_surjective K K _ (invFun_surjective e.injective) + rw [lift_umax.{u,v}, lift_id'.{u,v}] at this + have key := (lift_le.{v}.mpr <| max_aleph0_card_le_rank_fun_nat K).trans this + rw [lift_max, lift_aleph0, max_le_iff] at key + haveI : Infinite ιK := by + rw [← aleph0_le_mk_iff, bK.mk_eq_rank'']; exact key.1 + rw [bK.repr.toEquiv.cardinal_eq, mk_finsupp_lift_of_infinite, + lift_umax.{u,v}, lift_id'.{u,v}, bK.mk_eq_rank'', eq_comm, max_eq_left] + exact key.2 + +/-- The **Erdős-Kaplansky Theorem**: the dual of an infinite-dimensional vector space + over a division ring has dimension equal to its cardinality. -/ +theorem rank_dual_eq_card_dual_of_aleph0_le_rank' {V : Type*} [AddCommGroup V] [Module K V] + (h : ℵ₀ ≤ Module.rank K V) : Module.rank Kᵐᵒᵖ (V →ₗ[K] K) = #(V →ₗ[K] K) := by + obtain ⟨⟨ι, b⟩⟩ := Module.Free.exists_basis (R := K) (M := V) + rw [← b.mk_eq_rank'', aleph0_le_mk_iff] at h + have e := (b.constr Kᵐᵒᵖ (M' := K)).symm.trans + (LinearEquiv.piCongrRight fun _ ↦ MulOpposite.opLinearEquiv Kᵐᵒᵖ) + rw [e.rank_eq, e.toEquiv.cardinal_eq] + apply rank_fun_infinite + +/-- The **Erdős-Kaplansky Theorem** over a field. -/ +theorem rank_dual_eq_card_dual_of_aleph0_le_rank {K V} [Field K] [AddCommGroup V] [Module K V] + (h : ℵ₀ ≤ Module.rank K V) : Module.rank K (V →ₗ[K] K) = #(V →ₗ[K] K) := by + obtain ⟨⟨ι, b⟩⟩ := Module.Free.exists_basis (R := K) (M := V) + rw [← b.mk_eq_rank'', aleph0_le_mk_iff] at h + have e := (b.constr K (M' := K)).symm + rw [e.rank_eq, e.toEquiv.cardinal_eq] + apply rank_fun_infinite + +theorem lift_rank_lt_rank_dual' {V : Type v} [AddCommGroup V] [Module K V] + (h : ℵ₀ ≤ Module.rank K V) : + Cardinal.lift.{u} (Module.rank K V) < Module.rank Kᵐᵒᵖ (V →ₗ[K] K) := by + obtain ⟨⟨ι, b⟩⟩ := Module.Free.exists_basis (R := K) (M := V) + rw [← b.mk_eq_rank'', rank_dual_eq_card_dual_of_aleph0_le_rank' h, + ← (b.constr ℕ (M' := K)).toEquiv.cardinal_eq, mk_arrow] + apply cantor' + erw [nat_lt_lift_iff, one_lt_iff_nontrivial] + infer_instance + +theorem lift_rank_lt_rank_dual {K : Type u} {V : Type v} [Field K] [AddCommGroup V] [Module K V] + (h : ℵ₀ ≤ Module.rank K V) : + Cardinal.lift.{u} (Module.rank K V) < Module.rank K (V →ₗ[K] K) := by + rw [rank_dual_eq_card_dual_of_aleph0_le_rank h, ← rank_dual_eq_card_dual_of_aleph0_le_rank' h] + exact lift_rank_lt_rank_dual' h + +theorem rank_lt_rank_dual' {V : Type u} [AddCommGroup V] [Module K V] (h : ℵ₀ ≤ Module.rank K V) : + Module.rank K V < Module.rank Kᵐᵒᵖ (V →ₗ[K] K) := by + convert lift_rank_lt_rank_dual' h; rw [lift_id] + +theorem rank_lt_rank_dual {K V : Type u} [Field K] [AddCommGroup V] [Module K V] + (h : ℵ₀ ≤ Module.rank K V) : Module.rank K V < Module.rank K (V →ₗ[K] K) := by + convert lift_rank_lt_rank_dual h; rw [lift_id] + +end Cardinal diff --git a/Mathlib/LinearAlgebra/Dimension/Finite.lean b/Mathlib/LinearAlgebra/Dimension/Finite.lean index 3ae92933cf238..d764634e02e3d 100644 --- a/Mathlib/LinearAlgebra/Dimension/Finite.lean +++ b/Mathlib/LinearAlgebra/Dimension/Finite.lean @@ -3,7 +3,6 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Johannes Hölzl, Sander Dahmen, Kim Morrison -/ -import Mathlib.Algebra.Module.Torsion import Mathlib.SetTheory.Cardinal.Cofinality import Mathlib.LinearAlgebra.FreeModule.Finite.Basic import Mathlib.LinearAlgebra.Dimension.StrongRankCondition @@ -16,7 +15,6 @@ Also contains characterization for when rank equals zero or rank equals one. -/ - noncomputable section universe u v v' w @@ -101,11 +99,6 @@ theorem rank_pos [Nontrivial M] : 0 < Module.rank R M := end -lemma rank_eq_zero_iff_isTorsion {R M} [CommRing R] [IsDomain R] [AddCommGroup M] [Module R M] : - Module.rank R M = 0 ↔ Module.IsTorsion R M := by - rw [Module.IsTorsion, rank_eq_zero_iff] - simp [mem_nonZeroDivisors_iff_ne_zero] - variable (R M) /-- See `rank_subsingleton` that assumes `Subsingleton R` instead. -/ @@ -431,13 +424,6 @@ theorem Module.finrank_eq_zero_iff : rw [← rank_eq_zero_iff (R := R), ← finrank_eq_rank] norm_cast -/-- The `StrongRankCondition` is automatic. See `commRing_strongRankCondition`. -/ -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] - norm_cast - /-- A finite dimensional space has zero `finrank` iff it is a subsingleton. This is the `finrank` version of `rank_zero_iff`. -/ theorem Module.finrank_zero_iff [NoZeroSMulDivisors R M] : diff --git a/Mathlib/LinearAlgebra/Dimension/Torsion.lean b/Mathlib/LinearAlgebra/Dimension/Torsion/Basic.lean similarity index 100% rename from Mathlib/LinearAlgebra/Dimension/Torsion.lean rename to Mathlib/LinearAlgebra/Dimension/Torsion/Basic.lean diff --git a/Mathlib/LinearAlgebra/Dimension/Torsion/Finite.lean b/Mathlib/LinearAlgebra/Dimension/Torsion/Finite.lean new file mode 100644 index 0000000000000..6929e94e571f5 --- /dev/null +++ b/Mathlib/LinearAlgebra/Dimension/Torsion/Finite.lean @@ -0,0 +1,39 @@ +/- +Copyright (c) 2018 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro, Johannes Hölzl, Sander Dahmen, Kim Morrison +-/ +import Mathlib.Algebra.Module.Torsion +import Mathlib.LinearAlgebra.Dimension.Finite + +/-! +# Results relating rank and torsion. + +-/ + +universe u v + +variable {R : Type u} {M : Type v} +variable [Ring R] [AddCommGroup M] + +section RankZero + +variable [Nontrivial R] + +lemma rank_eq_zero_iff_isTorsion {R M} [CommRing R] [IsDomain R] [AddCommGroup M] [Module R M] : + Module.rank R M = 0 ↔ Module.IsTorsion R M := by + rw [Module.IsTorsion, rank_eq_zero_iff] + simp [mem_nonZeroDivisors_iff_ne_zero] + +end RankZero + +section StrongRankCondition + +/-- The `StrongRankCondition` is automatic. See `commRing_strongRankCondition`. -/ +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] + norm_cast + +end StrongRankCondition diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index 3870b51c3d7f0..4afc22536b4e1 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -11,6 +11,7 @@ import Mathlib.LinearAlgebra.SesquilinearForm import Mathlib.RingTheory.Finiteness.Projective import Mathlib.RingTheory.LocalRing.Basic import Mathlib.RingTheory.TensorProduct.Basic +import Mathlib.LinearAlgebra.Dimension.ErdosKaplansky /-! # Dual vector spaces @@ -1895,4 +1896,4 @@ noncomputable def dualDistribEquiv : Dual R M ⊗[R] Dual R N ≃ₗ[R] Dual R ( end TensorProduct -set_option linter.style.longFile 1900 +set_option linter.style.longFile 2100 diff --git a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean index 3003ced8d18b7..5aeea273a0985 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean @@ -7,7 +7,10 @@ import Mathlib.Algebra.Algebra.Spectrum import Mathlib.Algebra.Module.LinearMap.Basic import Mathlib.LinearAlgebra.GeneralLinearGroup import Mathlib.LinearAlgebra.FiniteDimensional +import Mathlib.RingTheory.Nilpotent.Defs +import Mathlib.RingTheory.Nilpotent.Lemmas import Mathlib.RingTheory.Nilpotent.Basic +import Mathlib.Tactic.Peel /-! # Eigenvectors and eigenvalues diff --git a/Mathlib/LinearAlgebra/FiniteDimensional.lean b/Mathlib/LinearAlgebra/FiniteDimensional.lean index 3b375180e50b1..0c769e5fc9c0b 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional.lean @@ -6,6 +6,7 @@ Authors: Chris Hughes import Mathlib.LinearAlgebra.FiniteDimensional.Defs import Mathlib.LinearAlgebra.Dimension.FreeAndStrongRankCondition import Mathlib.LinearAlgebra.Dimension.DivisionRing +import Mathlib.Tactic.IntervalCases /-! # Finite dimensional vector spaces @@ -18,6 +19,9 @@ Definitions and results that require fewer imports are in -/ +assert_not_exists Monoid.exponent +assert_not_exists Module.IsTorsion + universe u v v' diff --git a/Mathlib/LinearAlgebra/FreeModule/Finite/Matrix.lean b/Mathlib/LinearAlgebra/FreeModule/Finite/Matrix.lean index 218d9eb918e6a..bc49290326666 100644 --- a/Mathlib/LinearAlgebra/FreeModule/Finite/Matrix.lean +++ b/Mathlib/LinearAlgebra/FreeModule/Finite/Matrix.lean @@ -5,6 +5,7 @@ Authors: Riccardo Brasca -/ import Mathlib.LinearAlgebra.Dimension.LinearMap import Mathlib.LinearAlgebra.FreeModule.StrongRankCondition +import Mathlib.LinearAlgebra.Matrix.ToLin /-! # Finite and free modules using matrices diff --git a/Mathlib/LinearAlgebra/FreeModule/Norm.lean b/Mathlib/LinearAlgebra/FreeModule/Norm.lean index fb5d577bae4c0..29317643488c0 100644 --- a/Mathlib/LinearAlgebra/FreeModule/Norm.lean +++ b/Mathlib/LinearAlgebra/FreeModule/Norm.lean @@ -6,6 +6,7 @@ Authors: Junyan Xu import Mathlib.LinearAlgebra.FreeModule.IdealQuotient import Mathlib.RingTheory.Norm.Defs import Mathlib.RingTheory.AdjoinRoot +import Mathlib.LinearAlgebra.Dual /-! # Norms on free modules over principal ideal domains diff --git a/Mathlib/LinearAlgebra/Matrix/Diagonal.lean b/Mathlib/LinearAlgebra/Matrix/Diagonal.lean index ff4b403c74f6f..882fe25b80d27 100644 --- a/Mathlib/LinearAlgebra/Matrix/Diagonal.lean +++ b/Mathlib/LinearAlgebra/Matrix/Diagonal.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Patrick Massot, Casper Putz, Anne Baanen -/ import Mathlib.LinearAlgebra.Dimension.LinearMap +import Mathlib.LinearAlgebra.Matrix.ToLin /-! # Diagonal matrices diff --git a/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean b/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean index 51f9cb372d6b2..340f764a28faf 100644 --- a/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean +++ b/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean @@ -7,6 +7,7 @@ import Mathlib.Data.Matrix.Invertible import Mathlib.LinearAlgebra.FiniteDimensional.Defs import Mathlib.LinearAlgebra.Matrix.Adjugate import Mathlib.LinearAlgebra.Matrix.Trace +import Mathlib.LinearAlgebra.Matrix.ToLin /-! # Nonsingular inverses diff --git a/Mathlib/LinearAlgebra/Reflection.lean b/Mathlib/LinearAlgebra/Reflection.lean index 79d13d900e5b4..73d856bbbc9eb 100644 --- a/Mathlib/LinearAlgebra/Reflection.lean +++ b/Mathlib/LinearAlgebra/Reflection.lean @@ -8,6 +8,7 @@ import Mathlib.GroupTheory.OrderOfElement import Mathlib.LinearAlgebra.Dual import Mathlib.LinearAlgebra.FiniteSpan import Mathlib.RingTheory.Polynomial.Chebyshev +import Mathlib.Algebra.Module.Torsion /-! # Reflections in linear algebra diff --git a/Mathlib/LinearAlgebra/Semisimple.lean b/Mathlib/LinearAlgebra/Semisimple.lean index 34d9925e6aa6a..9f91ee1b94849 100644 --- a/Mathlib/LinearAlgebra/Semisimple.lean +++ b/Mathlib/LinearAlgebra/Semisimple.lean @@ -9,6 +9,7 @@ import Mathlib.LinearAlgebra.Basis.VectorSpace import Mathlib.RingTheory.Artinian import Mathlib.RingTheory.Ideal.Quotient.Nilpotent import Mathlib.RingTheory.SimpleModule +import Mathlib.Algebra.Module.Torsion /-! # Semisimple linear endomorphisms diff --git a/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean b/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean index 2c3bc40927bd3..e0b6251d52956 100644 --- a/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean +++ b/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean @@ -6,6 +6,7 @@ Authors: Sébastien Gouëzel import Mathlib.MeasureTheory.Measure.Lebesgue.EqHaar import Mathlib.MeasureTheory.Covering.Besicovitch import Mathlib.Tactic.AdaptationNote +import Mathlib.Algebra.EuclideanDomain.Basic /-! # Satellite configurations for Besicovitch covering lemma in vector spaces diff --git a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean index 0a214921062a9..67c398ceac956 100644 --- a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean @@ -6,6 +6,7 @@ Authors: Yury Kudryashov, Patrick Massot, Sébastien Gouëzel import Mathlib.Order.Interval.Set.Disjoint import Mathlib.MeasureTheory.Integral.SetIntegral import Mathlib.MeasureTheory.Measure.Lebesgue.Basic +import Mathlib.Algebra.EuclideanDomain.Basic /-! # Integral over an interval diff --git a/Mathlib/NumberTheory/NumberField/Units/Basic.lean b/Mathlib/NumberTheory/NumberField/Units/Basic.lean index 7907293a0c950..cecab8573fac6 100644 --- a/Mathlib/NumberTheory/NumberField/Units/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/Units/Basic.lean @@ -5,6 +5,7 @@ Authors: Xavier Roblot -/ import Mathlib.NumberTheory.NumberField.Embeddings import Mathlib.RingTheory.LocalRing.RingHom.Basic +import Mathlib.GroupTheory.Torsion /-! # Units of a number field diff --git a/Mathlib/RepresentationTheory/Basic.lean b/Mathlib/RepresentationTheory/Basic.lean index f17cfb213b7e5..7240eabf32ae6 100644 --- a/Mathlib/RepresentationTheory/Basic.lean +++ b/Mathlib/RepresentationTheory/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Antoine Labelle -/ import Mathlib.LinearAlgebra.Contraction +import Mathlib.Algebra.Group.Equiv.TypeTags /-! # Monoid representations diff --git a/Mathlib/RingTheory/Artinian.lean b/Mathlib/RingTheory/Artinian.lean index 2911d10aa8bb9..6bf257f5cca5f 100644 --- a/Mathlib/RingTheory/Artinian.lean +++ b/Mathlib/RingTheory/Artinian.lean @@ -11,6 +11,8 @@ import Mathlib.RingTheory.Nakayama import Mathlib.RingTheory.SimpleModule import Mathlib.Tactic.RSuffices import Mathlib.Tactic.StacksAttribute +import Mathlib.RingTheory.LocalRing.Basic +import Mathlib.RingTheory.Nilpotent.Lemmas /-! # Artinian rings and modules diff --git a/Mathlib/RingTheory/DedekindDomain/Ideal.lean b/Mathlib/RingTheory/DedekindDomain/Ideal.lean index 5925c3017eae3..8ad308eb008d5 100644 --- a/Mathlib/RingTheory/DedekindDomain/Ideal.lean +++ b/Mathlib/RingTheory/DedekindDomain/Ideal.lean @@ -9,6 +9,7 @@ import Mathlib.RingTheory.MaximalSpectrum import Mathlib.RingTheory.ChainOfDivisors import Mathlib.RingTheory.DedekindDomain.Basic import Mathlib.RingTheory.FractionalIdeal.Operations +import Mathlib.Algebra.Squarefree.Basic /-! # Dedekind domains and ideals diff --git a/Mathlib/RingTheory/LocalRing/Module.lean b/Mathlib/RingTheory/LocalRing/Module.lean index 6239fe6f8a670..40aa2ebe664c3 100644 --- a/Mathlib/RingTheory/LocalRing/Module.lean +++ b/Mathlib/RingTheory/LocalRing/Module.lean @@ -4,12 +4,12 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ import Mathlib.Algebra.Module.FinitePresentation -import Mathlib.LinearAlgebra.FiniteDimensional +import Mathlib.LinearAlgebra.Dual import Mathlib.RingTheory.FiniteType import Mathlib.RingTheory.Flat.Basic import Mathlib.RingTheory.LocalRing.ResidueField.Basic import Mathlib.RingTheory.Nakayama -import Mathlib.RingTheory.TensorProduct.Free +import Mathlib.Algebra.Module.Torsion /-! # Finite modules over local rings diff --git a/Mathlib/RingTheory/Polynomial/GaussLemma.lean b/Mathlib/RingTheory/Polynomial/GaussLemma.lean index 5766fbc7a9398..3af13f8f046bd 100644 --- a/Mathlib/RingTheory/Polynomial/GaussLemma.lean +++ b/Mathlib/RingTheory/Polynomial/GaussLemma.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ import Mathlib.FieldTheory.SplittingField.Construction -import Mathlib.RingTheory.Int.Basic import Mathlib.RingTheory.Localization.Integral import Mathlib.RingTheory.IntegralClosure.IntegrallyClosed import Mathlib.RingTheory.Polynomial.Content diff --git a/Mathlib/RingTheory/PowerBasis.lean b/Mathlib/RingTheory/PowerBasis.lean index a6ef6f783cb75..1e67541ed7447 100644 --- a/Mathlib/RingTheory/PowerBasis.lean +++ b/Mathlib/RingTheory/PowerBasis.lean @@ -5,6 +5,7 @@ Authors: Anne Baanen -/ import Mathlib.FieldTheory.Minpoly.Field import Mathlib.LinearAlgebra.SModEq +import Mathlib.RingTheory.Ideal.BigOperators /-! # Power basis From d1c574bfee9ae0a7c88fbfcdea799b3beb6dce84 Mon Sep 17 00:00:00 2001 From: Vasily Nesterov Date: Mon, 9 Dec 2024 13:17:13 +0000 Subject: [PATCH 594/829] refactor(Analysis/ODE): weaken Lipschitz condition (#19831) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Replace the assumption `∀ t, LipschitzOnWith K (v t) (s t)` with `∀ t ∈ I, LipschitzOnWith K (v t) (s t)` with an appropriate interval `I` in lemmas from `Analysis.ODE.Gronwall`. --- Mathlib/Analysis/ODE/Gronwall.lean | 48 +++++++++++-------- .../Manifold/IntegralCurve/ExistUnique.lean | 2 +- 2 files changed, 28 insertions(+), 22 deletions(-) diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index dd27e37df8326..011ee6ec7c7d1 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -127,9 +127,7 @@ theorem norm_le_gronwallBound_of_norm_deriv_right_le {f f' : ℝ → E} {δ K ε (fun x hx _r hr => (hf' x hx).liminf_right_slope_norm_le hr) ha bound variable {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} {f g f' g' : ℝ → E} {a b t₀ : ℝ} {εf εg δ : ℝ} - (hv : ∀ t, LipschitzOnWith K (v t) (s t)) -include hv in /-- If `f` and `g` are two approximate solutions of the same ODE, then the distance between them can't grow faster than exponentially. This is a simple corollary of Grönwall's inequality, and some people call this Grönwall's inequality too. @@ -137,6 +135,7 @@ people call this Grönwall's inequality too. This version assumes all inequalities to be true in some time-dependent set `s t`, and assumes that the solutions never leave this set. -/ theorem dist_le_of_approx_trajectories_ODE_of_mem + (hv : ∀ t ∈ Ico a b, LipschitzOnWith K (v t) (s t)) (hf : ContinuousOn f (Icc a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (f' t) (Ici t) t) (f_bound : ∀ t ∈ Ico a b, dist (f' t) (v t (f t)) ≤ εf) @@ -153,7 +152,7 @@ theorem dist_le_of_approx_trajectories_ODE_of_mem apply norm_le_gronwallBound_of_norm_deriv_right_le (hf.sub hg) h_deriv ha intro t ht have := dist_triangle4_right (f' t) (g' t) (v t (f t)) (v t (g t)) - have hv := (hv t).dist_le_mul _ (hfs t ht) _ (hgs t ht) + have hv := (hv t ht).dist_le_mul _ (hfs t ht) _ (hgs t ht) rw [← dist_eq_norm, ← dist_eq_norm] refine this.trans ((add_le_add (add_le_add (f_bound t ht) (g_bound t ht)) hv).trans ?_) rw [add_comm] @@ -174,10 +173,9 @@ theorem dist_le_of_approx_trajectories_ODE (ha : dist (f a) (g a) ≤ δ) : ∀ t ∈ Icc a b, dist (f t) (g t) ≤ gronwallBound δ K (εf + εg) (t - a) := have hfs : ∀ t ∈ Ico a b, f t ∈ @univ E := fun _ _ => trivial - dist_le_of_approx_trajectories_ODE_of_mem (fun t => (hv t).lipschitzOnWith) hf hf' + dist_le_of_approx_trajectories_ODE_of_mem (fun t _ => (hv t).lipschitzOnWith) hf hf' f_bound hfs hg hg' g_bound (fun _ _ => trivial) ha -include hv in /-- If `f` and `g` are two exact solutions of the same ODE, then the distance between them can't grow faster than exponentially. This is a simple corollary of Grönwall's inequality, and some people call this Grönwall's inequality too. @@ -185,6 +183,7 @@ people call this Grönwall's inequality too. This version assumes all inequalities to be true in some time-dependent set `s t`, and assumes that the solutions never leave this set. -/ theorem dist_le_of_trajectories_ODE_of_mem + (hv : ∀ t ∈ Ico a b, LipschitzOnWith K (v t) (s t)) (hf : ContinuousOn f (Icc a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hfs : ∀ t ∈ Ico a b, f t ∈ s t) @@ -212,16 +211,16 @@ theorem dist_le_of_trajectories_ODE (ha : dist (f a) (g a) ≤ δ) : ∀ t ∈ Icc a b, dist (f t) (g t) ≤ δ * exp (K * (t - a)) := have hfs : ∀ t ∈ Ico a b, f t ∈ @univ E := fun _ _ => trivial - dist_le_of_trajectories_ODE_of_mem (fun t => (hv t).lipschitzOnWith) hf hf' hfs hg + dist_le_of_trajectories_ODE_of_mem (fun t _ => (hv t).lipschitzOnWith) hf hf' hfs hg hg' (fun _ _ => trivial) ha -include hv in /-- There exists only one solution of an ODE \(\dot x=v(t, x)\) in a set `s ⊆ ℝ × E` with a given initial value provided that the RHS is Lipschitz continuous in `x` within `s`, and we consider only solutions included in `s`. This version shows uniqueness in a closed interval `Icc a b`, where `a` is the initial time. -/ theorem ODE_solution_unique_of_mem_Icc_right + (hv : ∀ t ∈ Ico a b, LipschitzOnWith K (v t) (s t)) (hf : ContinuousOn f (Icc a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hfs : ∀ t ∈ Ico a b, f t ∈ s t) @@ -233,10 +232,10 @@ theorem ODE_solution_unique_of_mem_Icc_right have := dist_le_of_trajectories_ODE_of_mem hv hf hf' hfs hg hg' hgs (dist_le_zero.2 ha) t ht rwa [zero_mul, dist_le_zero] at this -include hv in /-- A time-reversed version of `ODE_solution_unique_of_mem_Icc_right`. Uniqueness is shown in a closed interval `Icc a b`, where `b` is the "initial" time. -/ theorem ODE_solution_unique_of_mem_Icc_left + (hv : ∀ t ∈ Ioc a b, LipschitzOnWith K (v t) (s t)) (hf : ContinuousOn f (Icc a b)) (hf' : ∀ t ∈ Ioc a b, HasDerivWithinAt f (v t (f t)) (Iic t) t) (hfs : ∀ t ∈ Ioc a b, f t ∈ s t) @@ -245,9 +244,13 @@ theorem ODE_solution_unique_of_mem_Icc_left (hgs : ∀ t ∈ Ioc a b, g t ∈ s t) (hb : f b = g b) : EqOn f g (Icc a b) := by - have hv' t : LipschitzOnWith K (Neg.neg ∘ (v (-t))) (s (-t)) := by + have hv' : ∀ t ∈ Ico (-b) (-a), LipschitzOnWith K (Neg.neg ∘ (v (-t))) (s (-t)) := by + intro t ht + replace ht : -t ∈ Ioc a b := by + simp at ht ⊢ + constructor <;> linarith rw [← one_mul K] - exact LipschitzWith.id.neg.comp_lipschitzOnWith (hv _) + exact LipschitzWith.id.neg.comp_lipschitzOnWith (hv _ ht) have hmt1 : MapsTo Neg.neg (Icc (-b) (-a)) (Icc a b) := fun _ ht ↦ ⟨le_neg.mp ht.2, neg_le.mp ht.1⟩ have hmt2 : MapsTo Neg.neg (Ico (-b) (-a)) (Ioc a b) := @@ -270,10 +273,10 @@ theorem ODE_solution_unique_of_mem_Icc_left (hasDerivAt_neg t).hasDerivWithinAt (hmt3 t) simp -include hv in /-- A version of `ODE_solution_unique_of_mem_Icc_right` for uniqueness in a closed interval whose interior contains the initial time. -/ theorem ODE_solution_unique_of_mem_Icc + (hv : ∀ t ∈ Ioo a b, LipschitzOnWith K (v t) (s t)) (ht : t₀ ∈ Ioo a b) (hf : ContinuousOn f (Icc a b)) (hf' : ∀ t ∈ Ioo a b, HasDerivAt f (v t (f t)) t) @@ -286,21 +289,21 @@ theorem ODE_solution_unique_of_mem_Icc rw [← Icc_union_Icc_eq_Icc (le_of_lt ht.1) (le_of_lt ht.2)] apply EqOn.union · have hss : Ioc a t₀ ⊆ Ioo a b := Ioc_subset_Ioo_right ht.2 - exact ODE_solution_unique_of_mem_Icc_left hv + exact ODE_solution_unique_of_mem_Icc_left (fun t ht ↦ hv t (hss ht)) (hf.mono <| Icc_subset_Icc_right <| le_of_lt ht.2) (fun _ ht' ↦ (hf' _ (hss ht')).hasDerivWithinAt) (fun _ ht' ↦ (hfs _ (hss ht'))) (hg.mono <| Icc_subset_Icc_right <| le_of_lt ht.2) (fun _ ht' ↦ (hg' _ (hss ht')).hasDerivWithinAt) (fun _ ht' ↦ (hgs _ (hss ht'))) heq · have hss : Ico t₀ b ⊆ Ioo a b := Ico_subset_Ioo_left ht.1 - exact ODE_solution_unique_of_mem_Icc_right hv + exact ODE_solution_unique_of_mem_Icc_right (fun t ht ↦ hv t (hss ht)) (hf.mono <| Icc_subset_Icc_left <| le_of_lt ht.1) (fun _ ht' ↦ (hf' _ (hss ht')).hasDerivWithinAt) (fun _ ht' ↦ (hfs _ (hss ht'))) (hg.mono <| Icc_subset_Icc_left <| le_of_lt ht.1) (fun _ ht' ↦ (hg' _ (hss ht')).hasDerivWithinAt) (fun _ ht' ↦ (hgs _ (hss ht'))) heq -include hv in /-- A version of `ODE_solution_unique_of_mem_Icc` for uniqueness in an open interval. -/ theorem ODE_solution_unique_of_mem_Ioo + (hv : ∀ t ∈ Ioo a b, LipschitzOnWith K (v t) (s t)) (ht : t₀ ∈ Ioo a b) (hf : ∀ t ∈ Ioo a b, HasDerivAt f (v t (f t)) t ∧ f t ∈ s t) (hg : ∀ t ∈ Ioo a b, HasDerivAt g (v t (g t)) t ∧ g t ∈ s t) @@ -310,7 +313,8 @@ theorem ODE_solution_unique_of_mem_Ioo rcases lt_or_le t' t₀ with (h | h) · have hss : Icc t' t₀ ⊆ Ioo a b := fun _ ht'' ↦ ⟨lt_of_lt_of_le ht'.1 ht''.1, lt_of_le_of_lt ht''.2 ht.2⟩ - exact ODE_solution_unique_of_mem_Icc_left hv + exact ODE_solution_unique_of_mem_Icc_left + (fun t'' ht'' ↦ hv t'' ((Ioc_subset_Icc_self.trans hss) ht'')) (continuousOn_of_forall_continuousAt fun _ ht'' ↦ (hf _ <| hss ht'').1.continuousAt) (fun _ ht'' ↦ (hf _ <| hss <| Ioc_subset_Icc_self ht'').1.hasDerivWithinAt) (fun _ ht'' ↦ (hf _ <| hss <| Ioc_subset_Icc_self ht'').2) @@ -320,7 +324,8 @@ theorem ODE_solution_unique_of_mem_Ioo ⟨le_rfl, le_of_lt h⟩ · have hss : Icc t₀ t' ⊆ Ioo a b := fun _ ht'' ↦ ⟨lt_of_lt_of_le ht.1 ht''.1, lt_of_le_of_lt ht''.2 ht'.2⟩ - exact ODE_solution_unique_of_mem_Icc_right hv + exact ODE_solution_unique_of_mem_Icc_right + (fun t'' ht'' ↦ hv t'' ((Ico_subset_Icc_self.trans hss) ht'')) (continuousOn_of_forall_continuousAt fun _ ht'' ↦ (hf _ <| hss ht'').1.continuousAt) (fun _ ht'' ↦ (hf _ <| hss <| Ico_subset_Icc_self ht'').1.hasDerivWithinAt) (fun _ ht'' ↦ (hf _ <| hss <| Ico_subset_Icc_self ht'').2) @@ -329,18 +334,19 @@ theorem ODE_solution_unique_of_mem_Ioo (fun _ ht'' ↦ (hg _ <| hss <| Ico_subset_Icc_self ht'').2) heq ⟨h, le_rfl⟩ -include hv in /-- Local unqueness of ODE solutions. -/ theorem ODE_solution_unique_of_eventually + (hv : ∀ᶠ t in 𝓝 t₀, LipschitzOnWith K (v t) (s t)) (hf : ∀ᶠ t in 𝓝 t₀, HasDerivAt f (v t (f t)) t ∧ f t ∈ s t) (hg : ∀ᶠ t in 𝓝 t₀, HasDerivAt g (v t (g t)) t ∧ g t ∈ s t) (heq : f t₀ = g t₀) : f =ᶠ[𝓝 t₀] g := by - obtain ⟨ε, hε, h⟩ := eventually_nhds_iff_ball.mp (hf.and hg) + obtain ⟨ε, hε, h⟩ := eventually_nhds_iff_ball.mp (hv.and (hf.and hg)) rw [Filter.eventuallyEq_iff_exists_mem] refine ⟨ball t₀ ε, ball_mem_nhds _ hε, ?_⟩ simp_rw [Real.ball_eq_Ioo] at * - apply ODE_solution_unique_of_mem_Ioo hv (Real.ball_eq_Ioo t₀ ε ▸ mem_ball_self hε) - (fun _ ht ↦ (h _ ht).1) (fun _ ht ↦ (h _ ht).2) heq + apply ODE_solution_unique_of_mem_Ioo (fun _ ht ↦ (h _ ht).1) + (Real.ball_eq_Ioo t₀ ε ▸ mem_ball_self hε) + (fun _ ht ↦ (h _ ht).2.1) (fun _ ht ↦ (h _ ht).2.2) heq /-- There exists only one solution of an ODE \(\dot x=v(t, x)\) with a given initial value provided that the RHS is Lipschitz continuous in `x`. -/ @@ -353,5 +359,5 @@ theorem ODE_solution_unique (ha : f a = g a) : EqOn f g (Icc a b) := have hfs : ∀ t ∈ Ico a b, f t ∈ @univ E := fun _ _ => trivial - ODE_solution_unique_of_mem_Icc_right (fun t => (hv t).lipschitzOnWith) hf hf' hfs hg hg' + ODE_solution_unique_of_mem_Icc_right (fun t _ => (hv t).lipschitzOnWith) hf hf' hfs hg hg' (fun _ _ => trivial) ha diff --git a/Mathlib/Geometry/Manifold/IntegralCurve/ExistUnique.lean b/Mathlib/Geometry/Manifold/IntegralCurve/ExistUnique.lean index 8960304376097..bb586ece1fb8e 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve/ExistUnique.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve/ExistUnique.lean @@ -162,7 +162,7 @@ theorem isIntegralCurveAt_eventuallyEq_of_contMDiffAt (hγt₀ : I.IsInteriorPoi -- main proof suffices (extChartAt I (γ t₀)) ∘ γ =ᶠ[𝓝 t₀] (extChartAt I (γ' t₀)) ∘ γ' from (heq hγ).trans <| (this.fun_comp (extChartAt I (γ t₀)).symm).trans (h ▸ (heq hγ').symm) - exact ODE_solution_unique_of_eventually hlip + exact ODE_solution_unique_of_eventually (.of_forall hlip) (hdrv hγ rfl) (hdrv hγ' h) (by rw [Function.comp_apply, Function.comp_apply, h]) theorem isIntegralCurveAt_eventuallyEq_of_contMDiffAt_boundaryless [BoundarylessManifold I M] From 34fcc11f195e6652b791c9db2429dbca39abf5cb Mon Sep 17 00:00:00 2001 From: FMLJohn <417963616@qq.com> Date: Mon, 9 Dec 2024 13:17:15 +0000 Subject: [PATCH 595/829] doc(RingTheory/Polynomial/HilbertPoly): fix a typo (#19833) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In this pull request, we fix a typo in the file RingTheory/Polynomial/HilbertPoly.lean: originally, the docstring at the beginning of the file said "the polynomial `(n + 1)···(n + d)/d!`". But because the polynomial we mention is in `F[X]`, we must change it to `(X + 1)···(X + d)/d!`. --- Mathlib/RingTheory/Polynomial/HilbertPoly.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/RingTheory/Polynomial/HilbertPoly.lean b/Mathlib/RingTheory/Polynomial/HilbertPoly.lean index 806c1a591b0d2..66c5c57ec3060 100644 --- a/Mathlib/RingTheory/Polynomial/HilbertPoly.lean +++ b/Mathlib/RingTheory/Polynomial/HilbertPoly.lean @@ -21,10 +21,10 @@ This `h` is unique and is denoted as `Polynomial.hilbertPoly p d`. For example, given `d : ℕ`, the power series expansion of `1/(1-X)ᵈ⁺¹` in `F[X]` is `Σₙ ((d + n).choose d)Xⁿ`, which equals `Σₙ ((n + 1)···(n + d)/d!)Xⁿ` and hence -`Polynomial.hilbertPoly (1 : F[X]) (d + 1)` is the polynomial `(n + 1)···(n + d)/d!`. Note that -if `d! = 0` in `F`, then the polynomial `(n + 1)···(n + d)/d!` no longer works, so we do not -want the characteristic of `F` to be divisible by `d!`. As `Polynomial.hilbertPoly` may take -any `p : F[X]` and `d : ℕ` as its inputs, it is necessary for us to assume that `CharZero F`. +`Polynomial.hilbertPoly (1 : F[X]) (d + 1)` is the polynomial `(X + 1)···(X + d)/d!`. Note that +if `d! = 0` in `F`, then the polynomial `(X + 1)···(X + d)/d!` no longer works, so we do not want +the characteristic of `F` to be divisible by `d!`. As `Polynomial.hilbertPoly` may take any +`p : F[X]` and `d : ℕ` as its inputs, it is necessary for us to assume that `CharZero F`. ## Main definitions From f0c97b08aff6771cdf709a7fef23fe4b47232243 Mon Sep 17 00:00:00 2001 From: JovanGerb Date: Mon, 9 Dec 2024 13:17:16 +0000 Subject: [PATCH 596/829] chore: minimize import (#19835) It seems that this unused import isn't detected by shake --- Mathlib/Analysis/NormedSpace/HahnBanach/SeparatingDual.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/Analysis/NormedSpace/HahnBanach/SeparatingDual.lean b/Mathlib/Analysis/NormedSpace/HahnBanach/SeparatingDual.lean index 7684f1deeb4fc..ed0229bf183df 100644 --- a/Mathlib/Analysis/NormedSpace/HahnBanach/SeparatingDual.lean +++ b/Mathlib/Analysis/NormedSpace/HahnBanach/SeparatingDual.lean @@ -5,8 +5,9 @@ Authors: Sébastien Gouëzel -/ import Mathlib.Analysis.NormedSpace.HahnBanach.Extension import Mathlib.Analysis.NormedSpace.HahnBanach.Separation +import Mathlib.Analysis.NormedSpace.Multilinear.Basic +import Mathlib.Analysis.NormedSpace.OperatorNorm.Completeness import Mathlib.LinearAlgebra.Dual -import Mathlib.Analysis.Normed.Operator.BoundedLinearMaps /-! # Spaces with separating dual From b202b867947571f7d316d1f591be7e759dc0165c Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Mon, 9 Dec 2024 13:51:37 +0000 Subject: [PATCH 597/829] feat: non-degeneracy of root pairing form without assuming ordered scalars (#19767) The motivation, and headline item is the new lemma `RootPairing.rootForm_restrict_nondegenerate_of_isAnisotropic` which establishes non-degeneracy of the root form without assuming ordered scalars, at the cost of requiring the crystallographic assumption and field coefficients. Since we also want to keep the old lemma (renamed to) `RootPairing.rootForm_restrict_nondegenerate_of_ordered` which establishes the same result over an ordered ring, we need to do some work in order to unify things as much as possible. With this in mind I have: - Turned `RootPairing.IsCrystallographic` into a typeclass - Introduced `RootPairing.IsAnisotropic` (which I've just made up but seems to be a good way to unify things; I also considered the idea that the pairing should take values in an ordered subring but decided this was inferior) In order to prove the new results, a certain amount of API for bilinear forms needed to be added. This also enabled me to golf some of the proofs of the existing non-degeneracy result (for order scalars). I also included a small amount of golfing / moving, and to standardise on `LinearMap.ker P.RootForm` rather than `LinearMap.ker P.Polarisation` (since it interfaced better with the bilinear form API). As compensation I added the convenience lemma `RootPairing.ker_polarization_eq_ker_rootForm`. Co-authored-by: Oliver Nash <7734364+ocfnash@users.noreply.github.com> --- Mathlib/Algebra/Lie/Weights/RootSystem.lean | 6 +- .../BilinearForm/Orthogonal.lean | 8 +- .../LinearAlgebra/Dimension/RankNullity.lean | 7 +- Mathlib/LinearAlgebra/Dual.lean | 12 +- .../Eigenspace/Triangularizable.lean | 2 +- Mathlib/LinearAlgebra/FiniteDimensional.lean | 21 ++- Mathlib/LinearAlgebra/RootSystem/Defs.lean | 53 ++++-- .../RootSystem/Finite/CanonicalBilinear.lean | 71 ++++---- .../RootSystem/Finite/Nondegenerate.lean | 162 ++++++++++++++---- Mathlib/LinearAlgebra/SesquilinearForm.lean | 81 ++++++++- 10 files changed, 308 insertions(+), 115 deletions(-) diff --git a/Mathlib/Algebra/Lie/Weights/RootSystem.lean b/Mathlib/Algebra/Lie/Weights/RootSystem.lean index d2a8a94210cd9..c6334632c472d 100644 --- a/Mathlib/Algebra/Lie/Weights/RootSystem.lean +++ b/Mathlib/Algebra/Lie/Weights/RootSystem.lean @@ -407,9 +407,9 @@ alias rootSystem_toLin_apply := rootSystem_toPerfectPairing_apply @[simp] lemma rootSystem_root_apply (α) : (rootSystem H).root α = α := rfl @[simp] lemma rootSystem_coroot_apply (α) : (rootSystem H).coroot α = coroot α := rfl -theorem isCrystallographic_rootSystem : (rootSystem H).IsCrystallographic := by - rintro α _ ⟨β, rfl⟩ - exact ⟨chainBotCoeff β.1 α.1 - chainTopCoeff β.1 α.1, by simp [apply_coroot_eq_cast β.1 α.1]⟩ +instance : (rootSystem H).IsCrystallographic where + exists_int α β := + ⟨chainBotCoeff β.1 α.1 - chainTopCoeff β.1 α.1, by simp [apply_coroot_eq_cast β.1 α.1]⟩ theorem isReduced_rootSystem : (rootSystem H).IsReduced := by intro ⟨α, hα⟩ ⟨β, hβ⟩ e diff --git a/Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean b/Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean index bff969f2c534c..c003e3f2279c9 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean @@ -148,7 +148,11 @@ theorem orthogonal_le (h : N ≤ L) : B.orthogonal L ≤ B.orthogonal N := fun _ theorem le_orthogonal_orthogonal (b : B.IsRefl) : N ≤ B.orthogonal (B.orthogonal N) := fun n hn _ hm => b _ _ (hm n hn) -lemma orthogonal_top (hB : B.Nondegenerate) (hB₀ : B.IsRefl) : +lemma orthogonal_top_eq_ker (hB : B.IsRefl) : + B.orthogonal ⊤ = LinearMap.ker B := by + ext; simp [LinearMap.BilinForm.IsOrtho, LinearMap.ext_iff, hB.eq_iff] + +lemma orthogonal_top_eq_bot (hB : B.Nondegenerate) (hB₀ : B.IsRefl) : B.orthogonal ⊤ = ⊥ := (Submodule.eq_bot_iff _).mpr fun _ hx ↦ hB _ fun y ↦ hB₀ _ _ <| hx y Submodule.mem_top @@ -320,7 +324,7 @@ theorem finrank_add_finrank_orthogonal (b₁ : B.IsRefl) (W : Submodule K V) : lemma finrank_orthogonal (hB : B.Nondegenerate) (hB₀ : B.IsRefl) (W : Submodule K V) : finrank K (B.orthogonal W) = finrank K V - finrank K W := by have := finrank_add_finrank_orthogonal hB₀ (W := W) - rw [B.orthogonal_top hB hB₀, inf_bot_eq, finrank_bot, add_zero] at this + rw [B.orthogonal_top_eq_bot hB hB₀, inf_bot_eq, finrank_bot, add_zero] at this have : finrank K W ≤ finrank K V := finrank_le W omega diff --git a/Mathlib/LinearAlgebra/Dimension/RankNullity.lean b/Mathlib/LinearAlgebra/Dimension/RankNullity.lean index cf3c16a79daa0..48fd4277e7c4f 100644 --- a/Mathlib/LinearAlgebra/Dimension/RankNullity.lean +++ b/Mathlib/LinearAlgebra/Dimension/RankNullity.lean @@ -209,16 +209,17 @@ lemma Submodule.finrank_quotient [Module.Finite R M] {S : Type*} [Ring S] [SMul rw [← (N.restrictScalars R).finrank_quotient_add_finrank] exact Nat.eq_sub_of_add_eq rfl -lemma Submodule.disjoint_ker_of_finrank_eq [NoZeroSMulDivisors R M] {N : Type*} [AddCommGroup N] +lemma Submodule.disjoint_ker_of_finrank_le [NoZeroSMulDivisors R M] {N : Type*} [AddCommGroup N] [Module R N] {L : Submodule R M} [Module.Finite R L] (f : M →ₗ[R] N) - (h : finrank R (L.map f) = finrank R L) : + (h : finrank R L ≤ finrank R (L.map f)) : Disjoint L (LinearMap.ker f) := by refine disjoint_iff.mpr <| LinearMap.injective_domRestrict_iff.mp <| LinearMap.ker_eq_bot.mp <| Submodule.rank_eq_zero.mp ?_ rw [← Submodule.finrank_eq_rank, Nat.cast_eq_zero] rw [← LinearMap.range_domRestrict] at h have := (LinearMap.ker (f.domRestrict L)).finrank_quotient_add_finrank - rwa [LinearEquiv.finrank_eq (f.domRestrict L).quotKerEquivRange, h, Nat.add_eq_left] at this + rw [LinearEquiv.finrank_eq (f.domRestrict L).quotKerEquivRange] at this + omega end Finrank diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index 4afc22536b4e1..18b5e4ddd8cd1 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -1247,6 +1247,11 @@ noncomputable def quotEquivAnnihilator (W : Subspace K V) : (V ⧸ W) ≃ₗ[K] open Module +theorem finrank_add_finrank_dualAnnihilator_eq (W : Subspace K V) : + finrank K W + finrank K W.dualAnnihilator = finrank K V := by + rw [← W.quotEquivAnnihilator.finrank_eq (M₂ := dualAnnihilator W), + add_comm, Submodule.finrank_quotient_add_finrank] + @[simp] theorem finrank_dualCoannihilator_eq {Φ : Subspace K (Module.Dual K V)} : finrank K Φ.dualCoannihilator = finrank K Φ.dualAnnihilator := by @@ -1255,12 +1260,7 @@ theorem finrank_dualCoannihilator_eq {Φ : Subspace K (Module.Dual K V)} : theorem finrank_add_finrank_dualCoannihilator_eq (W : Subspace K (Module.Dual K V)) : finrank K W + finrank K W.dualCoannihilator = finrank K V := by - rw [finrank_dualCoannihilator_eq] - -- Porting note: LinearEquiv.finrank_eq needs help - let equiv := W.quotEquivAnnihilator - have eq := LinearEquiv.finrank_eq (R := K) (M := (Module.Dual K V) ⧸ W) - (M₂ := { x // x ∈ dualAnnihilator W }) equiv - rw [eq.symm, add_comm, Submodule.finrank_quotient_add_finrank, Subspace.dual_finrank_eq] + rw [finrank_dualCoannihilator_eq, finrank_add_finrank_dualAnnihilator_eq, dual_finrank_eq] end diff --git a/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean b/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean index 76112cc06a755..6a90c71055221 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean @@ -134,7 +134,7 @@ theorem iSup_maxGenEigenspace_eq_top [IsAlgClosed K] [FiniteDimensional K V] (f -- Since the dimensions of `ER` and `ES` add up to the dimension of `V`, it follows that the -- span of all generalized eigenvectors is all of `V`. show ⨆ (μ : K), f.maxGenEigenspace μ = ⊤ - rw [← top_le_iff, ← Submodule.eq_top_of_disjoint ER ES h_dim_add h_disjoint] + rw [← top_le_iff, ← Submodule.eq_top_of_disjoint ER ES h_dim_add.ge h_disjoint] apply sup_le hER hES -- Lemma 8.21 of [axler2015] diff --git a/Mathlib/LinearAlgebra/FiniteDimensional.lean b/Mathlib/LinearAlgebra/FiniteDimensional.lean index 0c769e5fc9c0b..3795257b53d17 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional.lean @@ -58,22 +58,29 @@ theorem finrank_add_le_finrank_add_finrank (s t : Submodule K V) [FiniteDimensio rw [← finrank_sup_add_finrank_inf_eq] exact self_le_add_right _ _ +theorem finrank_add_finrank_le_of_disjoint [FiniteDimensional K V] + {s t : Submodule K V} (hdisjoint : Disjoint s t) : + finrank K s + finrank K t ≤ finrank K V := by + rw [← Submodule.finrank_sup_add_finrank_inf_eq s t, hdisjoint.eq_bot, finrank_bot, add_zero] + exact Submodule.finrank_le _ + theorem eq_top_of_disjoint [FiniteDimensional K V] (s t : Submodule K V) - (hdim : finrank K s + finrank K t = finrank K V) (hdisjoint : Disjoint s t) : s ⊔ t = ⊤ := by + (hdim : finrank K V ≤ finrank K s + finrank K t) (hdisjoint : Disjoint s t) : s ⊔ t = ⊤ := by have h_finrank_inf : finrank K ↑(s ⊓ t) = 0 := by rw [disjoint_iff_inf_le, le_bot_iff] at hdisjoint rw [hdisjoint, finrank_bot] apply eq_top_of_finrank_eq - rw [← hdim] + replace hdim : finrank K V = finrank K s + finrank K t := + le_antisymm hdim (finrank_add_finrank_le_of_disjoint hdisjoint) + rw [hdim] convert s.finrank_sup_add_finrank_inf_eq t rw [h_finrank_inf] rfl -theorem finrank_add_finrank_le_of_disjoint [FiniteDimensional K V] - {s t : Submodule K V} (hdisjoint : Disjoint s t) : - finrank K s + finrank K t ≤ finrank K V := by - rw [← Submodule.finrank_sup_add_finrank_inf_eq s t, hdisjoint.eq_bot, finrank_bot, add_zero] - exact Submodule.finrank_le _ +theorem isCompl_iff_disjoint [FiniteDimensional K V] (s t : Submodule K V) + (hdim : finrank K V ≤ finrank K s + finrank K t) : + IsCompl s t ↔ Disjoint s t := + ⟨fun h ↦ h.1, fun h ↦ ⟨h, codisjoint_iff.mpr <| eq_top_of_disjoint s t hdim h⟩⟩ end DivisionRing diff --git a/Mathlib/LinearAlgebra/RootSystem/Defs.lean b/Mathlib/LinearAlgebra/RootSystem/Defs.lean index 66e76b9408a4a..da02ee54a1daf 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Defs.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Defs.lean @@ -353,21 +353,20 @@ lemma pairing_reflection_perm_self_right (i j : ι) : /-- 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.root' i) (range P.coroot) (zmultiples (1 : R)) +class IsCrystallographic : Prop where + exists_int : ∀ i j, ∃ z : ℤ, z = P.pairing i j + +protected lemma exists_int [P.IsCrystallographic] (i j : ι) : + ∃ z : ℤ, z = P.pairing i j := + IsCrystallographic.exists_int i j lemma isCrystallographic_iff : - P.IsCrystallographic ↔ ∀ i j, ∃ z : ℤ, z = P.pairing i j := by - rw [IsCrystallographic] - refine ⟨fun h i j ↦ ?_, fun h i _ ⟨j, hj⟩ ↦ ?_⟩ - · simpa [AddSubgroup.mem_zmultiples_iff] using h i (mem_range_self j) - · simpa [← hj, AddSubgroup.mem_zmultiples_iff] using h i j - -variable {P} in -lemma IsCrystallographic.flip (h : P.IsCrystallographic) : - P.flip.IsCrystallographic := by + P.IsCrystallographic ↔ ∀ i j, ∃ z : ℤ, z = P.pairing i j := + ⟨fun ⟨h⟩ ↦ h, fun h ↦ ⟨h⟩⟩ + +instance [P.IsCrystallographic] : P.flip.IsCrystallographic := by rw [isCrystallographic_iff, forall_comm] - exact P.isCrystallographic_iff.mp h + exact P.exists_int /-- A root pairing is said to be reduced if any linearly dependent pair of roots is related by a sign. -/ @@ -390,6 +389,30 @@ abbrev rootSpan := span R (range P.root) /-- The linear span of coroots. -/ abbrev corootSpan := span R (range P.coroot) +lemma coe_rootSpan_dualAnnihilator_map : + P.rootSpan.dualAnnihilator.map P.toDualRight.symm = {x | ∀ i, P.root' i x = 0} := by + ext x + rw [rootSpan, Submodule.map_coe, Submodule.coe_dualAnnihilator_span] + change x ∈ P.toDualRight.toEquiv.symm '' _ ↔ _ + rw [← Equiv.setOf_apply_symm_eq_image_setOf, Equiv.symm_symm] + simp [Set.range_subset_iff] + +lemma coe_corootSpan_dualAnnihilator_map : + P.corootSpan.dualAnnihilator.map P.toDualLeft.symm = {x | ∀ i, P.coroot' i x = 0} := + P.flip.coe_rootSpan_dualAnnihilator_map + +lemma rootSpan_dualAnnihilator_map_eq : + P.rootSpan.dualAnnihilator.map P.toDualRight.symm = + (span R (range P.root')).dualCoannihilator := by + apply SetLike.coe_injective + rw [Submodule.coe_dualCoannihilator_span, coe_rootSpan_dualAnnihilator_map] + simp + +lemma corootSpan_dualAnnihilator_map_eq : + P.corootSpan.dualAnnihilator.map P.toDualLeft.symm = + (span R (range P.coroot')).dualCoannihilator := + P.flip.rootSpan_dualAnnihilator_map_eq + /-- The `Weyl group` of a root pairing is the group of automorphisms of the weight space generated by reflections in roots. -/ def weylGroup : Subgroup (M ≃ₗ[R] M) := @@ -541,10 +564,10 @@ def coxeterWeight : R := pairing P i j * pairing P j i lemma coxeterWeight_swap : coxeterWeight P i j = coxeterWeight P j i := by simp only [coxeterWeight, mul_comm] -lemma exists_int_eq_coxeterWeight (h : P.IsCrystallographic) (i j : ι) : +lemma exists_int_eq_coxeterWeight [P.IsCrystallographic] (i j : ι) : ∃ z : ℤ, P.coxeterWeight i j = z := by - obtain ⟨a, ha⟩ := P.isCrystallographic_iff.mp h i j - obtain ⟨b, hb⟩ := P.isCrystallographic_iff.mp h j i + obtain ⟨a, ha⟩ := P.exists_int i j + obtain ⟨b, hb⟩ := P.exists_int j i exact ⟨a * b, by simp [coxeterWeight, ha, hb]⟩ /-- Two roots are orthogonal when they are fixed by each others' reflections. -/ diff --git a/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean b/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean index 3b1ebfefe84f2..0a28708c322c3 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean @@ -126,6 +126,14 @@ lemma coPolarization_apply_eq_zero_iff (n : N) : P.CoPolarization n = 0 ↔ P.CorootForm n = 0 := P.flip.polarization_apply_eq_zero_iff n +lemma ker_polarization_eq_ker_rootForm : + LinearMap.ker P.Polarization = LinearMap.ker P.RootForm := by + ext; simp only [LinearMap.mem_ker, P.polarization_apply_eq_zero_iff] + +lemma ker_copolarization_eq_ker_corootForm : + LinearMap.ker P.CoPolarization = LinearMap.ker P.CorootForm := + P.flip.ker_polarization_eq_ker_rootForm + lemma rootForm_symmetric : LinearMap.IsSymm P.RootForm := by simp [LinearMap.IsSymm, mul_comm, rootForm_apply_apply] @@ -196,6 +204,18 @@ theorem range_polarization_domRestrict_le_span_coroot : use fun i => (P.toPerfectPairing x) (P.coroot i) simp +theorem corootSpan_dualAnnihilator_le_ker_rootForm : + P.corootSpan.dualAnnihilator.map P.toDualLeft.symm ≤ LinearMap.ker P.RootForm := by + rw [← SetLike.coe_subset_coe, coe_corootSpan_dualAnnihilator_map] + intro x hx + simp only [coroot', PerfectPairing.flip_apply_apply, mem_setOf_eq] at hx + ext y + simp [rootForm_apply_apply, hx] + +theorem rootSpan_dualAnnihilator_le_ker_rootForm : + P.rootSpan.dualAnnihilator.map P.toDualRight.symm ≤ LinearMap.ker P.CorootForm := + P.flip.corootSpan_dualAnnihilator_le_ker_rootForm + lemma prod_rootForm_smul_coroot_mem_range_domRestrict (i : ι) : (∏ a : ι, P.RootForm (P.root a) (P.root a)) • P.coroot i ∈ LinearMap.range (P.Polarization.domRestrict (P.rootSpan)) := by @@ -206,29 +226,6 @@ lemma prod_rootForm_smul_coroot_mem_range_domRestrict (i : ι) : use ⟨(c • 2 • P.root i), by aesop⟩ simp -section IsCrystallographic - -variable [CharZero R] (h : P.IsCrystallographic) (i : ι) -include h - -lemma rootForm_apply_root_self_ne_zero : - P.RootForm (P.root i) (P.root i) ≠ 0 := by - choose z hz using P.isCrystallographic_iff.mp h i - simp only [rootForm_apply_apply, PerfectPairing.flip_apply_apply, root_coroot_eq_pairing, ← hz] - suffices 0 < ∑ i, z i * z i by norm_cast; exact this.ne' - refine Finset.sum_pos' (fun i _ ↦ mul_self_nonneg (z i)) ⟨i, Finset.mem_univ i, ?_⟩ - have hzi : z i = 2 := by - specialize hz i - rw [pairing_same] at hz - norm_cast at hz - simp [hzi] - -lemma corootForm_apply_coroot_self_ne_zero : - P.CorootForm (P.coroot i) (P.coroot i) ≠ 0 := - P.flip.rootForm_apply_root_self_ne_zero h.flip i - -end IsCrystallographic - end CommRing section LinearOrderedCommRing @@ -239,20 +236,14 @@ variable [Fintype ι] [LinearOrderedCommRing R] [AddCommGroup M] [Module R M] [A theorem rootForm_self_non_neg (x : M) : 0 ≤ P.RootForm x x := IsSumSq.nonneg (P.rootForm_self_sum_of_squares x) -theorem rootForm_self_zero_iff (x : M) : - P.RootForm x x = 0 ↔ ∀ i, P.coroot' i x = 0 := by - simp only [rootForm_apply_apply, PerfectPairing.toLin_apply, LinearMap.coe_comp, comp_apply, - Polarization_apply, map_sum, map_smul, smul_eq_mul] - convert Finset.sum_mul_self_eq_zero_iff Finset.univ fun i => P.coroot' i x - simp +lemma rootForm_self_eq_zero_iff {x : M} : + P.RootForm x x = 0 ↔ x ∈ LinearMap.ker P.RootForm := + P.RootForm.apply_apply_same_eq_zero_iff P.rootForm_self_non_neg P.rootForm_symmetric -lemma rootForm_root_self_pos (j : ι) : - 0 < P.RootForm (P.root j) (P.root j) := by - simp only [LinearMap.coe_mk, AddHom.coe_mk, LinearMap.coe_comp, comp_apply, - rootForm_apply_apply, toLin_toPerfectPairing] - refine Finset.sum_pos' (fun i _ => (sq (P.pairing j i)) ▸ sq_nonneg (P.pairing j i)) ?_ - use j - simp +lemma rootForm_root_self_pos (i : ι) : + 0 < P.RootForm (P.root i) (P.root i) := by + simp only [rootForm_apply_apply] + exact Finset.sum_pos' (fun j _ ↦ mul_self_nonneg _) ⟨i, by simp⟩ /-- SGA3 XXI Prop. 2.3.1 -/ lemma coxeterWeight_le_four (i j : ι) : P.coxeterWeight i j ≤ 4 := by @@ -274,10 +265,10 @@ instance instIsRootPositiveRootForm : IsRootPositive P P.RootForm where symm := P.rootForm_symmetric apply_reflection_eq := P.rootForm_reflection_reflection_apply -lemma coxeterWeight_mem_set_of_isCrystallographic (i j : ι) (hP : P.IsCrystallographic) : +lemma coxeterWeight_mem_set_of_isCrystallographic (i j : ι) [P.IsCrystallographic] : P.coxeterWeight i j ∈ ({0, 1, 2, 3, 4} : Set R) := by obtain ⟨n, hcn⟩ : ∃ n : ℕ, P.coxeterWeight i j = n := by - obtain ⟨z, hz⟩ := P.exists_int_eq_coxeterWeight hP i j + obtain ⟨z, hz⟩ := P.exists_int_eq_coxeterWeight i j have hz₀ : 0 ≤ z := by simpa [hz] using P.coxeterWeight_non_neg P.RootForm i j obtain ⟨n, rfl⟩ := Int.eq_ofNat_of_zero_le hz₀ exact ⟨n, by simp [hz]⟩ @@ -286,10 +277,6 @@ lemma coxeterWeight_mem_set_of_isCrystallographic (i j : ι) (hP : P.IsCrystallo norm_cast at this ⊢ omega -lemma prod_rootForm_root_self_pos : - 0 < ∏ i, P.RootForm (P.root i) (P.root i) := - Finset.prod_pos fun i _ => rootForm_root_self_pos P i - end LinearOrderedCommRing end RootPairing diff --git a/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean b/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean index 65fb8e8be87a7..82b70c781f91a 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Carnahan -/ import Mathlib.LinearAlgebra.BilinearForm.Basic +import Mathlib.LinearAlgebra.BilinearForm.Orthogonal import Mathlib.LinearAlgebra.Dimension.Localization import Mathlib.LinearAlgebra.QuadraticForm.Basic import Mathlib.LinearAlgebra.RootSystem.Finite.CanonicalBilinear @@ -22,13 +23,16 @@ Another application is to the faithfulness of the Weyl group action on roots, an Weyl group. ## Main results: - * `RootPairing.rootForm_rootPositive`: `RootForm` is root-positive. - * `RootPairing.polarization_domRestrict_injective`: The polarization restricted to the span of - roots is injective. + * `RootPairing.IsAnisotropic`: We say a finite root pairing is anisotropic if there are no roots / + coroots which have length zero wrt the root / coroot forms. * `RootPairing.rootForm_pos_of_nonzero`: `RootForm` is strictly positive on non-zero linear combinations of roots. This gives us a convenient way to eliminate certain Dynkin diagrams from the classification, since it suffices to produce a nonzero linear combination of simple roots with non-positive norm. + * `RootPairing.rootForm_restrict_nondegenerate_of_ordered`: The root form is non-degenerate if + the coefficients are ordered. + * `RootPairing.rootForm_restrict_nondegenerate_of_isAnisotropic`: the root form is + non-degenerate if the coefficients are a field and the pairing is crystallographic. ## References: * [N. Bourbaki, *Lie groups and {L}ie algebras. {C}hapters 4--6*][bourbaki1968] @@ -49,10 +53,50 @@ open Submodule (span) namespace RootPairing -variable {ι R M N : Type*} +variable {ι R M N : Type*} [Fintype ι] [AddCommGroup M] [AddCommGroup N] -variable [Fintype ι] [LinearOrderedCommRing R] [AddCommGroup M] [Module R M] [AddCommGroup N] - [Module R N] (P : RootPairing ι R M N) +section CommRing + +variable [CommRing R] [Module R M] [Module R N] (P : RootPairing ι R M N) + +/-- We say a finite root pairing is anisotropic if there are no roots / coroots which have length +zero wrt the root / coroot forms. + +Examples include crystallographic pairings in characteristic zero +`RootPairing.instIsAnisotropicOfIsCrystallographic` and pairings over ordered scalars. +`RootPairing.instIsAnisotropicOfLinearOrderedCommRing`. -/ +class IsAnisotropic : Prop where + rootForm_root_ne_zero (i : ι) : P.RootForm (P.root i) (P.root i) ≠ 0 + corootForm_coroot_ne_zero (i : ι) : P.CorootForm (P.coroot i) (P.coroot i) ≠ 0 + +instance [P.IsAnisotropic] : P.flip.IsAnisotropic where + rootForm_root_ne_zero := IsAnisotropic.corootForm_coroot_ne_zero + corootForm_coroot_ne_zero := IsAnisotropic.rootForm_root_ne_zero + +/-- An auxiliary lemma en route to `RootPairing.instIsAnisotropicOfIsCrystallographic`. -/ +private lemma rootForm_root_ne_zero_aux [CharZero R] [P.IsCrystallographic] (i : ι) : + P.RootForm (P.root i) (P.root i) ≠ 0 := by + choose z hz using P.exists_int i + simp only [rootForm_apply_apply, PerfectPairing.flip_apply_apply, root_coroot_eq_pairing, ← hz] + suffices 0 < ∑ i, z i * z i by norm_cast; exact this.ne' + refine Finset.sum_pos' (fun i _ ↦ mul_self_nonneg (z i)) ⟨i, Finset.mem_univ i, ?_⟩ + have hzi : z i = 2 := by + specialize hz i + rw [pairing_same] at hz + norm_cast at hz + simp [hzi] + +instance instIsAnisotropicOfIsCrystallographic [CharZero R] [P.IsCrystallographic] : + IsAnisotropic P where + rootForm_root_ne_zero := P.rootForm_root_ne_zero_aux + corootForm_coroot_ne_zero := P.flip.rootForm_root_ne_zero_aux + +end CommRing + +section IsDomain + +variable [CommRing R] [IsDomain R] [Module R M] [Module R N] (P : RootPairing ι R M N) + [P.IsAnisotropic] @[simp] lemma finrank_rootSpan_map_polarization_eq_finrank_corootSpan : @@ -60,9 +104,11 @@ lemma finrank_rootSpan_map_polarization_eq_finrank_corootSpan : rw [← LinearMap.range_domRestrict] apply (Submodule.finrank_mono P.range_polarization_domRestrict_le_span_coroot).antisymm have : IsReflexive R N := PerfectPairing.reflexive_right P.toPerfectPairing + have h_ne : ∏ i, P.RootForm (P.root i) (P.root i) ≠ 0 := + Finset.prod_ne_zero_iff.mpr fun i _ ↦ IsAnisotropic.rootForm_root_ne_zero i refine LinearMap.finrank_le_of_isSMulRegular P.corootSpan (LinearMap.range (P.Polarization.domRestrict P.rootSpan)) - (smul_right_injective N (Ne.symm (ne_of_lt P.prod_rootForm_root_self_pos))) + (smul_right_injective N h_ne) fun _ hx => ?_ obtain ⟨c, hc⟩ := (mem_span_range_iff_exists_fun R).mp hx rw [← hc, Finset.smul_sum] @@ -80,39 +126,97 @@ lemma finrank_corootSpan_eq : finrank R P.corootSpan = finrank R P.rootSpan := le_antisymm P.finrank_corootSpan_le P.flip.finrank_corootSpan_le -lemma disjoint_rootSpan_ker_polarization : - Disjoint P.rootSpan (LinearMap.ker P.Polarization) := by +lemma disjoint_rootSpan_ker_rootForm : + Disjoint P.rootSpan (LinearMap.ker P.RootForm) := by have : IsReflexive R M := PerfectPairing.reflexive_left P.toPerfectPairing - refine Submodule.disjoint_ker_of_finrank_eq (L := P.rootSpan) P.Polarization ?_ - rw [finrank_rootSpan_map_polarization_eq_finrank_corootSpan, finrank_corootSpan_eq] + rw [← P.ker_polarization_eq_ker_rootForm] + refine Submodule.disjoint_ker_of_finrank_le (L := P.rootSpan) P.Polarization ?_ + rw [P.finrank_rootSpan_map_polarization_eq_finrank_corootSpan, P.finrank_corootSpan_eq] -lemma mem_ker_polarization_of_rootForm_self_eq_zero {x : M} (hx : P.RootForm x x = 0) : - x ∈ LinearMap.ker P.Polarization := by - rw [LinearMap.mem_ker, Polarization_apply] - rw [rootForm_self_zero_iff] at hx - exact Fintype.sum_eq_zero _ fun i ↦ by simp [hx i] +lemma disjoint_corootSpan_ker_corootForm : + Disjoint P.corootSpan (LinearMap.ker P.CorootForm) := + P.flip.disjoint_rootSpan_ker_rootForm + +end IsDomain + +section Field + +variable [Field R] [Module R M] [Module R N] (P : RootPairing ι R M N) [P.IsAnisotropic] + +lemma isCompl_rootSpan_ker_rootForm : + IsCompl P.rootSpan (LinearMap.ker P.RootForm) := by + have _iM : IsReflexive R M := PerfectPairing.reflexive_left P.toPerfectPairing + have _iN : IsReflexive R N := PerfectPairing.reflexive_right P.toPerfectPairing + refine (Submodule.isCompl_iff_disjoint _ _ ?_).mpr P.disjoint_rootSpan_ker_rootForm + have aux : finrank R M = finrank R P.rootSpan + finrank R P.corootSpan.dualAnnihilator := by + rw [P.toPerfectPairing.finrank_eq, ← P.finrank_corootSpan_eq, + Subspace.finrank_add_finrank_dualAnnihilator_eq P.corootSpan] + rw [aux, add_le_add_iff_left] + convert Submodule.finrank_mono P.corootSpan_dualAnnihilator_le_ker_rootForm + exact (LinearEquiv.finrank_map_eq _ _).symm + +lemma isCompl_corootSpan_ker_corootForm : + IsCompl P.corootSpan (LinearMap.ker P.CorootForm) := + P.flip.isCompl_rootSpan_ker_rootForm + +/-- See also `RootPairing.rootForm_restrict_nondegenerate_of_ordered`. + +Note that this applies to crystallographic root systems in characteristic zero via +`RootPairing.instIsAnisotropicOfIsCrystallographic`. -/ +lemma rootForm_restrict_nondegenerate_of_isAnisotropic : + LinearMap.Nondegenerate (P.RootForm.restrict P.rootSpan) := + P.rootForm_symmetric.nondegenerate_restrict_of_isCompl_ker P.isCompl_rootSpan_ker_rootForm + +@[simp] +lemma orthogonal_rootSpan_eq : + P.RootForm.orthogonal P.rootSpan = LinearMap.ker P.RootForm := by + rw [← LinearMap.BilinForm.orthogonal_top_eq_ker P.rootForm_symmetric.isRefl] + refine le_antisymm ?_ (by intro; aesop) + rintro x hx y - + simp only [LinearMap.BilinForm.mem_orthogonal_iff, LinearMap.BilinForm.IsOrtho] at hx ⊢ + obtain ⟨u, hu, v, hv, rfl⟩ : ∃ᵉ (u ∈ P.rootSpan) (v ∈ LinearMap.ker P.RootForm), u + v = y := by + rw [← Submodule.mem_sup, P.isCompl_rootSpan_ker_rootForm.sup_eq_top]; exact Submodule.mem_top + simp only [LinearMap.mem_ker] at hv + simp [hx _ hu, hv] + +@[simp] +lemma orthogonal_corootSpan_eq : + P.CorootForm.orthogonal P.corootSpan = LinearMap.ker P.CorootForm := + P.flip.orthogonal_rootSpan_eq + +end Field + +section LinearOrderedCommRing + +variable [LinearOrderedCommRing R] [Module R M] [Module R N] (P : RootPairing ι R M N) + +instance instIsAnisotropicOfLinearOrderedCommRing : IsAnisotropic P where + rootForm_root_ne_zero i := (P.rootForm_root_self_pos i).ne' + corootForm_coroot_ne_zero i := (P.flip.rootForm_root_self_pos i).ne' + +/-- See also `RootPairing.rootForm_restrict_nondegenerate_of_isAnisotropic`. -/ +lemma rootForm_restrict_nondegenerate_of_ordered : + LinearMap.Nondegenerate (P.RootForm.restrict P.rootSpan) := + (P.RootForm.nondegenerate_restrict_iff_disjoint_ker (rootForm_self_non_neg P) + P.rootForm_symmetric).mpr P.disjoint_rootSpan_ker_rootForm lemma eq_zero_of_mem_rootSpan_of_rootForm_self_eq_zero {x : M} (hx : x ∈ P.rootSpan) (hx' : P.RootForm x x = 0) : x = 0 := by - rw [← Submodule.mem_bot (R := R), ← P.disjoint_rootSpan_ker_polarization.eq_bot] - exact ⟨hx, P.mem_ker_polarization_of_rootForm_self_eq_zero hx'⟩ + have : x ∈ P.rootSpan ⊓ LinearMap.ker P.RootForm := ⟨hx, P.rootForm_self_eq_zero_iff.mp hx'⟩ + simpa [P.disjoint_rootSpan_ker_rootForm.eq_bot] using this + +lemma rootForm_pos_of_ne_zero {x : M} (hx : x ∈ P.rootSpan) (h : x ≠ 0) : + 0 < P.RootForm x x := by + apply (P.rootForm_self_non_neg x).lt_of_ne + contrapose! h + exact P.eq_zero_of_mem_rootSpan_of_rootForm_self_eq_zero hx h.symm lemma _root_.RootSystem.rootForm_anisotropic (P : RootSystem ι R M N) : P.RootForm.toQuadraticMap.Anisotropic := fun x ↦ P.eq_zero_of_mem_rootSpan_of_rootForm_self_eq_zero <| by simpa only [rootSpan, P.span_eq_top] using Submodule.mem_top -lemma rootForm_pos_of_nonzero {x : M} (hx : x ∈ P.rootSpan) (h : x ≠ 0) : - 0 < P.RootForm x x := by - apply (P.rootForm_self_non_neg x).lt_of_ne - contrapose! h - exact eq_zero_of_mem_rootSpan_of_rootForm_self_eq_zero P hx h.symm - -lemma rootForm_restrict_nondegenerate : - (P.RootForm.restrict P.rootSpan).Nondegenerate := - LinearMap.IsRefl.nondegenerate_of_separatingLeft (LinearMap.IsSymm.isRefl fun x y => by - simp [rootForm_apply_apply, mul_comm]) fun x h => SetLike.coe_eq_coe.mp - (P.eq_zero_of_mem_rootSpan_of_rootForm_self_eq_zero (Submodule.coe_mem x) (h x)) +end LinearOrderedCommRing end RootPairing diff --git a/Mathlib/LinearAlgebra/SesquilinearForm.lean b/Mathlib/LinearAlgebra/SesquilinearForm.lean index f4f55354935c4..495286b640a80 100644 --- a/Mathlib/LinearAlgebra/SesquilinearForm.lean +++ b/Mathlib/LinearAlgebra/SesquilinearForm.lean @@ -155,6 +155,8 @@ include H theorem eq_zero : ∀ {x y}, B x y = 0 → B y x = 0 := fun {x y} ↦ H x y +theorem eq_iff {x y} : B x y = 0 ↔ B y x = 0 := ⟨H x y, H y x⟩ + theorem ortho_comm {x y} : IsOrtho B x y ↔ IsOrtho B y x := ⟨eq_zero H, eq_zero H⟩ @@ -700,25 +702,54 @@ section CommRing variable [CommRing R] [AddCommGroup M] [Module R M] [AddCommGroup M₁] [Module R M₁] {I I' : R →+* R} -theorem IsRefl.nondegenerate_of_separatingLeft {B : M →ₗ[R] M →ₗ[R] M₁} (hB : B.IsRefl) - (hB' : B.SeparatingLeft) : B.Nondegenerate := by - refine ⟨hB', ?_⟩ +theorem IsRefl.nondegenerate_iff_separatingLeft {B : M →ₗ[R] M →ₗ[R] M₁} (hB : B.IsRefl) : + B.Nondegenerate ↔ B.SeparatingLeft := by + refine ⟨fun h ↦ h.1, fun hB' ↦ ⟨hB', ?_⟩⟩ rw [separatingRight_iff_flip_ker_eq_bot, hB.ker_eq_bot_iff_ker_flip_eq_bot.mp] rwa [← separatingLeft_iff_ker_eq_bot] -theorem IsRefl.nondegenerate_of_separatingRight {B : M →ₗ[R] M →ₗ[R] M₁} (hB : B.IsRefl) - (hB' : B.SeparatingRight) : B.Nondegenerate := by - refine ⟨?_, hB'⟩ +theorem IsRefl.nondegenerate_iff_separatingRight {B : M →ₗ[R] M →ₗ[R] M₁} (hB : B.IsRefl) : + B.Nondegenerate ↔ B.SeparatingRight := by + refine ⟨fun h ↦ h.2, fun hB' ↦ ⟨?_, hB'⟩⟩ rw [separatingLeft_iff_ker_eq_bot, hB.ker_eq_bot_iff_ker_flip_eq_bot.mpr] rwa [← separatingRight_iff_flip_ker_eq_bot] +lemma disjoint_ker_of_nondegenerate_restrict {B : M →ₗ[R] M →ₗ[R] M₁} {W : Submodule R M} + (hW : (B.domRestrict₁₂ W W).Nondegenerate) : + Disjoint W (LinearMap.ker B) := by + refine Submodule.disjoint_def.mpr fun x hx hx' ↦ ?_ + let x' : W := ⟨x, hx⟩ + suffices x' = 0 by simpa [x'] + apply hW.1 x' + simp_rw [Subtype.forall, domRestrict₁₂_apply] + intro y hy + rw [mem_ker] at hx' + simp [hx'] + +lemma IsSymm.nondegenerate_restrict_of_isCompl_ker {B : M →ₗ[R] M →ₗ[R] R} (hB : B.IsSymm) + {W : Submodule R M} (hW : IsCompl W (LinearMap.ker B)) : + (B.domRestrict₁₂ W W).Nondegenerate := by + have hB' : (B.domRestrict₁₂ W W).IsRefl := fun x y ↦ hB.isRefl (W.subtype x) (W.subtype y) + rw [LinearMap.IsRefl.nondegenerate_iff_separatingLeft hB'] + intro ⟨x, hx⟩ hx' + simp only [Submodule.mk_eq_zero] + replace hx' : ∀ y ∈ W, B x y = 0 := by simpa [Subtype.forall] using hx' + replace hx' : x ∈ W ⊓ ker B := by + refine ⟨hx, ?_⟩ + ext y + obtain ⟨u, hu, v, hv, rfl⟩ : ∃ u ∈ W, ∃ v ∈ ker B, u + v = y := by + rw [← Submodule.mem_sup, hW.sup_eq_top]; exact Submodule.mem_top + suffices B x u = 0 by rw [mem_ker] at hv; simpa [← hB.eq v, hv] + exact hx' u hu + simpa [hW.inf_eq_bot] using hx' + /-- The restriction of a reflexive bilinear map `B` onto a submodule `W` is nondegenerate if `W` has trivial intersection with its orthogonal complement, that is `Disjoint W (W.orthogonalBilin B)`. -/ theorem nondegenerate_restrict_of_disjoint_orthogonal {B : M →ₗ[R] M →ₗ[R] M₁} (hB : B.IsRefl) {W : Submodule R M} (hW : Disjoint W (W.orthogonalBilin B)) : (B.domRestrict₁₂ W W).Nondegenerate := by - refine (hB.domRestrict W).nondegenerate_of_separatingLeft ?_ + rw [(hB.domRestrict W).nondegenerate_iff_separatingLeft] rintro ⟨x, hx⟩ b₁ rw [Submodule.mk_eq_zero, ← Submodule.mem_bot R] refine hW.le_bot ⟨hx, fun y hy ↦ ?_⟩ @@ -889,6 +920,42 @@ lemma apply_sq_lt_iff_linearIndependent_of_symm [NoZeroSMulDivisors R M] rw [show (B x y) ^ 2 = (B x y) * (B y x) by rw [sq, ← hB, RingHom.id_apply]] exact apply_mul_apply_lt_iff_linearIndependent B hp x y +lemma apply_apply_same_eq_zero_iff (hs : ∀ x, 0 ≤ B x x) (hB : B.IsSymm) {x : M} : + B x x = 0 ↔ x ∈ LinearMap.ker B := by + rw [LinearMap.mem_ker] + refine ⟨fun h ↦ ?_, fun h ↦ by simp [h]⟩ + ext y + have := B.apply_sq_le_of_symm hs hB x y + simp only [h, zero_mul] at this + exact pow_eq_zero <| le_antisymm this (sq_nonneg (B x y)) + +lemma nondegenerate_iff (hs : ∀ x, 0 ≤ B x x) (hB : B.IsSymm) : + B.Nondegenerate ↔ ∀ x, B x x = 0 ↔ x = 0 := by + simp_rw [hB.isRefl.nondegenerate_iff_separatingLeft, separatingLeft_iff_ker_eq_bot, + Submodule.eq_bot_iff, B.apply_apply_same_eq_zero_iff hs hB, mem_ker] + exact forall_congr' fun x ↦ by aesop + +/-- A convenience variant of `LinearMap.BilinForm.nondegenerate_iff` characterising nondegeneracy as +positive definiteness. -/ +lemma nondegenerate_iff' (hs : ∀ x, 0 ≤ B x x) (hB : B.IsSymm) : + B.Nondegenerate ↔ ∀ x, x ≠ 0 → 0 < B x x := by + rw [B.nondegenerate_iff hs hB, ← not_iff_not] + push_neg + exact exists_congr fun x ↦ ⟨by aesop, fun ⟨h₀, h⟩ ↦ Or.inl ⟨le_antisymm h (hs x), h₀⟩⟩ + +lemma nondegenerate_restrict_iff_disjoint_ker (hs : ∀ x, 0 ≤ B x x) (hB : B.IsSymm) + {W : Submodule R M} : + (B.domRestrict₁₂ W W).Nondegenerate ↔ Disjoint W (LinearMap.ker B) := by + refine ⟨disjoint_ker_of_nondegenerate_restrict, fun hW ↦ ?_⟩ + have hB' : (B.domRestrict₁₂ W W).IsRefl := fun x y ↦ hB.isRefl (W.subtype x) (W.subtype y) + rw [IsRefl.nondegenerate_iff_separatingLeft hB'] + intro ⟨x, hx⟩ h + simp_rw [Subtype.forall, domRestrict₁₂_apply] at h + specialize h x hx + rw [B.apply_apply_same_eq_zero_iff hs hB] at h + have key : x ∈ W ⊓ LinearMap.ker B := ⟨hx, h⟩ + simpa [hW.eq_bot] using key + end BilinForm end LinearMap From 2c7acde63a9963ef7c316e2909716f510f8661e4 Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Mon, 9 Dec 2024 17:33:19 +0000 Subject: [PATCH 598/829] chore(Module/LocalizedModule): rename `ringHom_ext` to `ext` (#19787) `IsLocalizedModule.ringHom_ext` does not have any `RingHom` in its type signature. --- Mathlib/Algebra/Module/FinitePresentation.lean | 10 +++++----- Mathlib/Algebra/Module/LocalizedModule/Basic.lean | 9 ++++++--- Mathlib/RingTheory/LocalProperties/Projective.lean | 4 ++-- 3 files changed, 13 insertions(+), 10 deletions(-) diff --git a/Mathlib/Algebra/Module/FinitePresentation.lean b/Mathlib/Algebra/Module/FinitePresentation.lean index b5737ac1b8e49..8d27146b99787 100644 --- a/Mathlib/Algebra/Module/FinitePresentation.lean +++ b/Mathlib/Algebra/Module/FinitePresentation.lean @@ -373,7 +373,7 @@ lemma exists_bijective_map_powers [Module.Finite R M] [Module.FinitePresentation simp only [Module.algebraMap_end_apply, algebraMap_smul, LinearMap.map_smul_of_tower] rw [LinearMap.smul_comp, ← smul_assoc s₀.1, Algebra.smul_def s₀.1, IsUnit.mul_val_inv, one_smul] apply LinearMap.restrictScalars_injective R - apply IsLocalizedModule.ringHom_ext (.powers t) (LocalizedModule.mkLinearMap (.powers t) M) + apply IsLocalizedModule.ext (.powers t) (LocalizedModule.mkLinearMap (.powers t) M) (IsLocalizedModule.map_units (LocalizedModule.mkLinearMap (.powers t) M)) ext x have : s₂.1 • l' (l x) = s₂.1 • s₀.1 • x := congr($hs₂ x) @@ -384,7 +384,7 @@ lemma exists_bijective_map_powers [Module.Finite R M] [Module.FinitePresentation simp only [Module.algebraMap_end_apply, algebraMap_smul, LinearMap.map_smul_of_tower] rw [LinearMap.comp_smul, ← smul_assoc s₀.1, Algebra.smul_def s₀.1, IsUnit.mul_val_inv, one_smul] apply LinearMap.restrictScalars_injective R - apply IsLocalizedModule.ringHom_ext (.powers t) (LocalizedModule.mkLinearMap (.powers t) N) + apply IsLocalizedModule.ext (.powers t) (LocalizedModule.mkLinearMap (.powers t) N) (IsLocalizedModule.map_units (LocalizedModule.mkLinearMap (.powers t) N)) ext x have : s₁.1 • l (l' x) = s₁.1 • s₀.1 • x := congr($hs₁ x) @@ -415,7 +415,7 @@ lemma Module.FinitePresentation.exists_lift_equiv_of_isLocalizedModule obtain ⟨l', s, H⟩ := Module.FinitePresentation.exists_lift_of_isLocalizedModule S g (l ∘ₗ f) have : Function.Bijective (IsLocalizedModule.map S f g l') := by have : IsLocalizedModule.map S f g l' = (s • LinearMap.id) ∘ₗ l := by - apply IsLocalizedModule.ringHom_ext S f (IsLocalizedModule.map_units g) + apply IsLocalizedModule.ext S f (IsLocalizedModule.map_units g) apply LinearMap.ext fun x ↦ ?_ simp only [LinearMap.coe_comp, Function.comp_apply, IsLocalizedModule.map_apply, Basis.coe_repr_symm, LinearMap.coe_restrictScalars] @@ -433,7 +433,7 @@ lemma Module.FinitePresentation.exists_lift_equiv_of_isLocalizedModule ((Module.End_isUnit_iff _).mp ((hsu.unit⁻¹).isUnit.map (algebraMap _ (End Rᵣₛ (LocalizedModule rs N))))).comp (hr' (r * s) (dvd_mul_right _ _)) refine ⟨r * s, mul_mem hr s.2, LinearEquiv.ofBijective _ this, ?_⟩ - apply IsLocalizedModule.ringHom_ext rs (LocalizedModule.mkLinearMap rs M) fun x ↦ map_units g + apply IsLocalizedModule.ext rs (LocalizedModule.mkLinearMap rs M) fun x ↦ map_units g ⟨x.1, SetLike.le_def.mp (Submonoid.powers_le.mpr (mul_mem hr s.2)) x.2⟩ ext x apply ((Module.End_isUnit_iff _).mp (IsLocalizedModule.map_units g s)).1 @@ -461,7 +461,7 @@ instance Module.FinitePresentation.isLocalizedModule_map [Module.FinitePresentat · intro h obtain ⟨h', s, e⟩ := Module.FinitePresentation.exists_lift_of_isLocalizedModule S g (h ∘ₗ f) refine ⟨⟨h', s⟩, ?_⟩ - apply IsLocalizedModule.ringHom_ext S f (IsLocalizedModule.map_units g) + apply IsLocalizedModule.ext S f (IsLocalizedModule.map_units g) refine e.symm.trans (by ext; simp) · intro h₁ h₂ e apply Module.Finite.exists_smul_of_comp_eq_of_isLocalizedModule S g diff --git a/Mathlib/Algebra/Module/LocalizedModule/Basic.lean b/Mathlib/Algebra/Module/LocalizedModule/Basic.lean index b7f0ab08698d3..3cdf9bb43a356 100644 --- a/Mathlib/Algebra/Module/LocalizedModule/Basic.lean +++ b/Mathlib/Algebra/Module/LocalizedModule/Basic.lean @@ -919,11 +919,14 @@ theorem linearMap_ext {N N'} [AddCommMonoid N] [Module R N] [AddCommMonoid N'] [ (h : g ∘ₗ f = g' ∘ₗ f) : g = g' := (is_universal S f _ <| map_units f').unique h rfl -theorem ringHom_ext (map_unit : ∀ x : S, IsUnit ((algebraMap R (Module.End R M'')) x)) +theorem ext (map_unit : ∀ x : S, IsUnit ((algebraMap R (Module.End R M'')) x)) ⦃j k : M' →ₗ[R] M''⦄ (h : j.comp f = k.comp f) : j = k := by rw [← lift_unique S f (k.comp f) map_unit j h, lift_unique] rfl +@[deprecated (since := "2024-12-07")] +alias ringHom_ext := ext + /-- If `(M', f)` and `(M'', g)` both satisfy universal property of localized module, then `M', M''` are isomorphic as `R`-module -/ @@ -1115,10 +1118,10 @@ noncomputable def map : (M →ₗ[R] N) →ₗ[R] (M' →ₗ[R] N') where toFun h := lift S f (g ∘ₗ h) (IsLocalizedModule.map_units g) map_add' h₁ h₂ := by - apply IsLocalizedModule.ringHom_ext S f (IsLocalizedModule.map_units g) + apply IsLocalizedModule.ext S f (IsLocalizedModule.map_units g) simp only [lift_comp, LinearMap.add_comp, LinearMap.comp_add] map_smul' r h := by - apply IsLocalizedModule.ringHom_ext S f (IsLocalizedModule.map_units g) + apply IsLocalizedModule.ext S f (IsLocalizedModule.map_units g) simp only [lift_comp, LinearMap.add_comp, LinearMap.comp_add, LinearMap.smul_comp, LinearMap.comp_smul, RingHom.id_apply] diff --git a/Mathlib/RingTheory/LocalProperties/Projective.lean b/Mathlib/RingTheory/LocalProperties/Projective.lean index d6c0268663ed7..bf3d45d29deed 100644 --- a/Mathlib/RingTheory/LocalProperties/Projective.lean +++ b/Mathlib/RingTheory/LocalProperties/Projective.lean @@ -95,7 +95,7 @@ theorem LinearMap.split_surjective_of_localization_maximal conv_lhs => rw [← LinearMap.map_smul_of_tower] rw [← Submonoid.smul_def, IsLocalizedModule.mk'_cancel', IsLocalizedModule.mk'_cancel'] apply LinearMap.restrictScalars_injective R - apply IsLocalizedModule.ringHom_ext I.primeCompl (LocalizedModule.mkLinearMap I.primeCompl N) + apply IsLocalizedModule.ext I.primeCompl (LocalizedModule.mkLinearMap I.primeCompl N) · exact IsLocalizedModule.map_units (LocalizedModule.mkLinearMap I.primeCompl N) ext simp only [LocalizedModule.map_mk, LinearMap.coe_comp, LinearMap.coe_restrictScalars, @@ -111,7 +111,7 @@ theorem LinearMap.split_surjective_of_localization_maximal simp only [Module.algebraMap_end_apply, ← Submonoid.smul_def, IsLocalizedModule.mk'_cancel', ← LinearMap.map_smul_of_tower] apply LinearMap.restrictScalars_injective R - apply IsLocalizedModule.ringHom_ext I.primeCompl (LocalizedModule.mkLinearMap I.primeCompl N) + apply IsLocalizedModule.ext I.primeCompl (LocalizedModule.mkLinearMap I.primeCompl N) · exact IsLocalizedModule.map_units (LocalizedModule.mkLinearMap I.primeCompl N) ext simp only [coe_comp, coe_restrictScalars, Function.comp_apply, From 0ab5ec968d8ba4a05a82181af1f05f49fedf13de Mon Sep 17 00:00:00 2001 From: sgouezel Date: Mon, 9 Dec 2024 17:46:02 +0000 Subject: [PATCH 599/829] feat: differentiability of functions on manifolds built by applying linear maps to a section of a vector bundle (#19636) We already have the same kind of statement for higher smoothness, but not for mere differentiability. --- Mathlib.lean | 1 + .../VectorBundle/MDifferentiable.lean | 148 ++++++++++++++++++ 2 files changed, 149 insertions(+) create mode 100644 Mathlib/Geometry/Manifold/VectorBundle/MDifferentiable.lean diff --git a/Mathlib.lean b/Mathlib.lean index f6c3ebb239a3b..07fe8c7ea85df 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3098,6 +3098,7 @@ import Mathlib.Geometry.Manifold.SmoothManifoldWithCorners import Mathlib.Geometry.Manifold.VectorBundle.Basic import Mathlib.Geometry.Manifold.VectorBundle.FiberwiseLinear import Mathlib.Geometry.Manifold.VectorBundle.Hom +import Mathlib.Geometry.Manifold.VectorBundle.MDifferentiable import Mathlib.Geometry.Manifold.VectorBundle.Pullback import Mathlib.Geometry.Manifold.VectorBundle.SmoothSection import Mathlib.Geometry.Manifold.VectorBundle.Tangent diff --git a/Mathlib/Geometry/Manifold/VectorBundle/MDifferentiable.lean b/Mathlib/Geometry/Manifold/VectorBundle/MDifferentiable.lean new file mode 100644 index 0000000000000..3f9eae5d70e38 --- /dev/null +++ b/Mathlib/Geometry/Manifold/VectorBundle/MDifferentiable.lean @@ -0,0 +1,148 @@ +/- +Copyright (c) 2024 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel +-/ +import Mathlib.Geometry.Manifold.VectorBundle.Basic +import Mathlib.Geometry.Manifold.MFDeriv.NormedSpace +import Mathlib.Geometry.Manifold.MFDeriv.SpecificFunctions + + +/-! +# Differentiability of functions in vector bundles + +-/ + +open Bundle Set PartialHomeomorph ContinuousLinearMap Pretrivialization Filter + +open scoped Manifold Bundle Topology + +section + + +variable {𝕜 B B' F M : Type*} {E : B → Type*} + +variable [NontriviallyNormedField 𝕜] [NormedAddCommGroup F] [NormedSpace 𝕜 F] + [TopologicalSpace (TotalSpace F E)] [∀ x, TopologicalSpace (E x)] {EB : Type*} + [NormedAddCommGroup EB] [NormedSpace 𝕜 EB] {HB : Type*} [TopologicalSpace HB] + (IB : ModelWithCorners 𝕜 EB HB) (E' : B → Type*) [∀ x, Zero (E' x)] {EM : Type*} + [NormedAddCommGroup EM] [NormedSpace 𝕜 EM] {HM : Type*} [TopologicalSpace HM] + {IM : ModelWithCorners 𝕜 EM HM} [TopologicalSpace M] [ChartedSpace HM M] + {n : ℕ∞} + + +variable [TopologicalSpace B] [ChartedSpace HB B] [FiberBundle F E] + +/-- Characterization of differentiable functions into a smooth vector bundle. -/ +theorem mdifferentiableWithinAt_totalSpace (f : M → TotalSpace F E) {s : Set M} {x₀ : M} : + MDifferentiableWithinAt IM (IB.prod 𝓘(𝕜, F)) f s x₀ ↔ + MDifferentiableWithinAt IM IB (fun x => (f x).proj) s x₀ ∧ + MDifferentiableWithinAt IM 𝓘(𝕜, F) + (fun x ↦ (trivializationAt F E (f x₀).proj (f x)).2) s x₀ := by + simp (config := { singlePass := true }) only [mdifferentiableWithinAt_iff_target] + rw [and_and_and_comm, ← FiberBundle.continuousWithinAt_totalSpace, and_congr_right_iff] + intro hf + simp_rw [modelWithCornersSelf_prod, FiberBundle.extChartAt, Function.comp_def, + PartialEquiv.trans_apply, PartialEquiv.prod_coe, PartialEquiv.refl_coe, + extChartAt_self_apply, modelWithCornersSelf_coe, Function.id_def, ← chartedSpaceSelf_prod] + refine (mdifferentiableWithinAt_prod_iff _).trans (and_congr ?_ Iff.rfl) + have h1 : (fun x => (f x).proj) ⁻¹' (trivializationAt F E (f x₀).proj).baseSet ∈ 𝓝[s] x₀ := + ((FiberBundle.continuous_proj F E).continuousWithinAt.comp hf (mapsTo_image f s)) + ((Trivialization.open_baseSet _).mem_nhds (mem_baseSet_trivializationAt F E _)) + refine EventuallyEq.mdifferentiableWithinAt_iff (eventually_of_mem h1 fun x hx => ?_) ?_ + · simp_rw [Function.comp, PartialHomeomorph.coe_coe, Trivialization.coe_coe] + rw [Trivialization.coe_fst'] + exact hx + · simp only [mfld_simps] + +end + + +section + +/- Declare two manifolds `B₁` and `B₂` (with models `IB₁ : HB₁ → EB₁` and `IB₂ : HB₂ → EB₂`), +and two vector bundles `E₁` and `E₂` respectively over `B₁` and `B₂` (with model fibers +`F₁` and `F₂`). + +Also a third manifold `M`, which will be the source of all our maps. +-/ +variable {𝕜 F₁ F₂ B₁ B₂ M : Type*} {E₁ : B₁ → Type*} {E₂ : B₂ → Type*} [NontriviallyNormedField 𝕜] + [∀ x, AddCommGroup (E₁ x)] [∀ x, Module 𝕜 (E₁ x)] [NormedAddCommGroup F₁] [NormedSpace 𝕜 F₁] + [TopologicalSpace (TotalSpace F₁ E₁)] [∀ x, TopologicalSpace (E₁ x)] [∀ x, AddCommGroup (E₂ x)] + [∀ x, Module 𝕜 (E₂ x)] [NormedAddCommGroup F₂] [NormedSpace 𝕜 F₂] + [TopologicalSpace (TotalSpace F₂ E₂)] [∀ x, TopologicalSpace (E₂ x)] + {EB₁ : Type*} + [NormedAddCommGroup EB₁] [NormedSpace 𝕜 EB₁] {HB₁ : Type*} [TopologicalSpace HB₁] + {IB₁ : ModelWithCorners 𝕜 EB₁ HB₁} [TopologicalSpace B₁] [ChartedSpace HB₁ B₁] + {EB₂ : Type*} + [NormedAddCommGroup EB₂] [NormedSpace 𝕜 EB₂] {HB₂ : Type*} [TopologicalSpace HB₂] + {IB₂ : ModelWithCorners 𝕜 EB₂ HB₂} [TopologicalSpace B₂] [ChartedSpace HB₂ B₂] + {EM : Type*} + [NormedAddCommGroup EM] [NormedSpace 𝕜 EM] {HM : Type*} [TopologicalSpace HM] + {IM : ModelWithCorners 𝕜 EM HM} [TopologicalSpace M] [ChartedSpace HM M] + {n : ℕ∞} [FiberBundle F₁ E₁] [VectorBundle 𝕜 F₁ E₁] + [FiberBundle F₂ E₂] [VectorBundle 𝕜 F₂ E₂] + {b₁ : M → B₁} {b₂ : M → B₂} {m₀ : M} + {ϕ : Π (m : M), E₁ (b₁ m) →L[𝕜] E₂ (b₂ m)} {v : Π (m : M), E₁ (b₁ m)} {s : Set M} + +/-- Consider a smooth map `v : M → E₁` to a vector bundle, over a basemap `b₁ : M → B₁`, and +another basemap `b₂ : M → B₂`. Given linear maps `ϕ m : E₁ (b₁ m) → E₂ (b₂ m)` depending smoothly +on `m`, one can apply `ϕ m` to `g m`, and the resulting map is smooth. + +Note that the smoothness of `ϕ` can not be always be stated as smoothness of a map into a manifold, +as the pullback bundles `b₁ *ᵖ E₁` and `b₂ *ᵖ E₂` only make sense when `b₁` and `b₂` are globally +smooth, but we want to apply this lemma with only local information. Therefore, we formulate it +using smoothness of `ϕ` read in coordinates. + +Version for `MDifferentiableWithinAt`. We also give a version for `MDifferentiableAt`, but no +version for `MDifferentiableOn` or `MDifferentiable` as our assumption, written in coordinates, +only makes sense around a point. + -/ +lemma MDifferentiableWithinAt.clm_apply_of_inCoordinates + (hϕ : MDifferentiableWithinAt IM 𝓘(𝕜, F₁ →L[𝕜] F₂) + (fun m ↦ inCoordinates F₁ E₁ F₂ E₂ (b₁ m₀) (b₁ m) (b₂ m₀) (b₂ m) (ϕ m)) s m₀) + (hv : MDifferentiableWithinAt IM (IB₁.prod 𝓘(𝕜, F₁)) (fun m ↦ (v m : TotalSpace F₁ E₁)) s m₀) + (hb₂ : MDifferentiableWithinAt IM IB₂ b₂ s m₀) : + MDifferentiableWithinAt IM (IB₂.prod 𝓘(𝕜, F₂)) + (fun m ↦ (ϕ m (v m) : TotalSpace F₂ E₂)) s m₀ := by + rw [mdifferentiableWithinAt_totalSpace] at hv ⊢ + refine ⟨hb₂, ?_⟩ + apply (MDifferentiableWithinAt.clm_apply hϕ hv.2).congr_of_eventuallyEq_insert + have A : ∀ᶠ m in 𝓝[insert m₀ s] m₀, b₁ m ∈ (trivializationAt F₁ E₁ (b₁ m₀)).baseSet := by + apply hv.1.insert.continuousWithinAt + apply (trivializationAt F₁ E₁ (b₁ m₀)).open_baseSet.mem_nhds + exact FiberBundle.mem_baseSet_trivializationAt' (b₁ m₀) + have A' : ∀ᶠ m in 𝓝[insert m₀ s] m₀, b₂ m ∈ (trivializationAt F₂ E₂ (b₂ m₀)).baseSet := by + apply hb₂.insert.continuousWithinAt + apply (trivializationAt F₂ E₂ (b₂ m₀)).open_baseSet.mem_nhds + exact FiberBundle.mem_baseSet_trivializationAt' (b₂ m₀) + filter_upwards [A, A'] with m hm h'm + rw [inCoordinates_eq hm h'm] + simp only [coe_comp', ContinuousLinearEquiv.coe_coe, Trivialization.continuousLinearEquivAt_apply, + Trivialization.continuousLinearEquivAt_symm_apply, Function.comp_apply] + congr + rw [Trivialization.symm_apply_apply_mk (trivializationAt F₁ E₁ (b₁ m₀)) hm (v m)] + +/-- Consider a smooth map `v : M → E₁` to a vector bundle, over a basemap `b₁ : M → B₁`, and +another basemap `b₂ : M → B₂`. Given linear maps `ϕ m : E₁ (b₁ m) → E₂ (b₂ m)` depending smoothly +on `m`, one can apply `ϕ m` to `g m`, and the resulting map is smooth. + +Note that the smoothness of `ϕ` can not be always be stated as smoothness of a map into a manifold, +as the pullback bundles `b₁ *ᵖ E₁` and `b₂ *ᵖ E₂` only make sense when `b₁` and `b₂` are globally +smooth, but we want to apply this lemma with only local information. Therefore, we formulate it +using smoothness of `ϕ` read in coordinates. + +Version for `MDifferentiableAt`. We also give a version for `MDifferentiableWithinAt`, +but no version for `MDifferentiableOn` or `MDifferentiable` as our assumption, written +in coordinates, only makes sense around a point. + -/ +lemma MDifferentiableAt.clm_apply_of_inCoordinates + (hϕ : MDifferentiableAt IM 𝓘(𝕜, F₁ →L[𝕜] F₂) + (fun m ↦ inCoordinates F₁ E₁ F₂ E₂ (b₁ m₀) (b₁ m) (b₂ m₀) (b₂ m) (ϕ m)) m₀) + (hv : MDifferentiableAt IM (IB₁.prod 𝓘(𝕜, F₁)) (fun m ↦ (v m : TotalSpace F₁ E₁)) m₀) + (hb₂ : MDifferentiableAt IM IB₂ b₂ m₀) : + MDifferentiableAt IM (IB₂.prod 𝓘(𝕜, F₂)) (fun m ↦ (ϕ m (v m) : TotalSpace F₂ E₂)) m₀ := by + rw [← mdifferentiableWithinAt_univ] at hϕ hv hb₂ ⊢ + exact MDifferentiableWithinAt.clm_apply_of_inCoordinates hϕ hv hb₂ + +end From 3c0727b0ede04f28db6ce9f8fbca7071b62a0bc9 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Mon, 9 Dec 2024 17:46:04 +0000 Subject: [PATCH 600/829] feat: the homeomorphism between the tangent bundle to the model space and the product is a diffeomorphism (#19637) --- .../Geometry/Manifold/MFDeriv/Tangent.lean | 11 +++ .../Manifold/MFDeriv/UniqueDifferential.lean | 19 +++- .../Manifold/VectorBundle/Tangent.lean | 88 ++++++++++++++++++- 3 files changed, 113 insertions(+), 5 deletions(-) diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Tangent.lean b/Mathlib/Geometry/Manifold/MFDeriv/Tangent.lean index e3df71c1fd846..2c048b112606f 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Tangent.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Tangent.lean @@ -6,6 +6,7 @@ Authors: Sébastien Gouëzel, Floris van Doorn import Mathlib.Geometry.Manifold.MFDeriv.Atlas import Mathlib.Geometry.Manifold.MFDeriv.UniqueDifferential import Mathlib.Geometry.Manifold.VectorBundle.Tangent +import Mathlib.Geometry.Manifold.Diffeomorph /-! # Derivatives of maps in the tangent bundle @@ -87,3 +88,13 @@ lemma inTangentCoordinates_eq_mfderiv_comp · apply mdifferentiableWithinAt_extChartAt_symm apply (extChartAt I (f x₀)).map_source simpa using hx + +open Bundle +variable (I) in +/-- The canonical identification between the tangent bundle to the model space and the product, +as a diffeomorphism -/ +def tangentBundleModelSpaceDiffeomorph (n : ℕ∞) : + TangentBundle I H ≃ₘ^n⟮I.tangent, I.prod 𝓘(𝕜, E)⟯ ModelProd H E where + __ := TotalSpace.toProd H E + contMDiff_toFun := contMDiff_tangentBundleModelSpaceHomeomorph + contMDiff_invFun := contMDiff_tangentBundleModelSpaceHomeomorph_symm diff --git a/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean b/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean index 187e4cca99ec5..2382f7093e3f8 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean @@ -85,16 +85,29 @@ theorem UniqueMDiffOn.uniqueMDiffOn_preimage (hs : UniqueMDiffOn I s) {e : Parti variable [SmoothManifoldWithCorners I M] in /-- If a set in a manifold has the unique derivative property, then its pullback by any extended chart, in the vector space, also has the unique derivative property. -/ -theorem UniqueMDiffOn.uniqueDiffOn_target_inter (hs : UniqueMDiffOn I s) (x : M) : - UniqueDiffOn 𝕜 ((extChartAt I x).target ∩ (extChartAt I x).symm ⁻¹' s) := by +theorem UniqueMDiffOn.uniqueMDiffOn_target_inter (hs : UniqueMDiffOn I s) (x : M) : + UniqueMDiffOn 𝓘(𝕜, E) ((extChartAt I x).target ∩ (extChartAt I x).symm ⁻¹' s) := by -- this is just a reformulation of `UniqueMDiffOn.uniqueMDiffOn_preimage`, using as `e` -- the local chart at `x`. - apply UniqueMDiffOn.uniqueDiffOn rw [← PartialEquiv.image_source_inter_eq', inter_comm, extChartAt_source] exact (hs.inter (chartAt H x).open_source).image_denseRange' (fun y hy ↦ hasMFDerivWithinAt_extChartAt hy.2) fun y hy ↦ ((mdifferentiable_chart _).mfderiv_surjective hy.2).denseRange +variable [SmoothManifoldWithCorners I M] in +/-- If a set in a manifold has the unique derivative property, then its pullback by any extended +chart, in the vector space, also has the unique derivative property. -/ +theorem UniqueMDiffOn.uniqueDiffOn_target_inter (hs : UniqueMDiffOn I s) (x : M) : + UniqueDiffOn 𝕜 ((extChartAt I x).target ∩ (extChartAt I x).symm ⁻¹' s) := + (hs.uniqueMDiffOn_target_inter x).uniqueDiffOn + +variable [SmoothManifoldWithCorners I M] in +theorem UniqueMDiffOn.uniqueDiffWithinAt_range_inter (hs : UniqueMDiffOn I s) (x : M) (y : E) + (hy : y ∈ (extChartAt I x).target ∩ (extChartAt I x).symm ⁻¹' s) : + UniqueDiffWithinAt 𝕜 (range I ∩ (extChartAt I x).symm ⁻¹' s) y := by + apply (hs.uniqueDiffOn_target_inter x y hy).mono + apply inter_subset_inter_left _ (extChartAt_target_subset_range x) + variable [SmoothManifoldWithCorners I M] in /-- When considering functions between manifolds, this statement shows up often. It entails the unique differential of the pullback in extended charts of the set where the function can diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean index 5931fefe4f64b..05b06edfad0ea 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean @@ -300,6 +300,14 @@ end TangentBundleInstances /-! ## The tangent bundle to the model space -/ +@[simp, mfld_simps] +theorem trivializationAt_model_space_apply (p : TangentBundle I H) (x : H) : + trivializationAt E (TangentSpace I) x p = (p.1, p.2) := by + simp [TangentBundle.trivializationAt_apply] + have : fderivWithin 𝕜 (↑I ∘ ↑I.symm) (range I) (I p.proj) = + fderivWithin 𝕜 id (range I) (I p.proj) := + fderivWithin_congr' (fun y hy ↦ by simp [hy]) (mem_range_self p.proj) + simp [this, fderivWithin_id (ModelWithCorners.uniqueDiffWithinAt_image I)] /-- In the tangent bundle to the model space, the charts are just the canonical identification between a product type and a sigma type, a.k.a. `TotalSpace.toProd`. -/ @@ -336,7 +344,7 @@ theorem tangentBundleCore_coordChange_model_space (x x' z : H) : variable (I) in /-- The canonical identification between the tangent bundle to the model space and the product, -as a homeomorphism -/ +as a homeomorphism. For the diffeomorphism version, see `tangentBundleModelSpaceDiffeomorph`. -/ def tangentBundleModelSpaceHomeomorph : TangentBundle I H ≃ₜ ModelProd H E := { TotalSpace.toProd H E with continuous_toFun := by @@ -366,11 +374,83 @@ theorem tangentBundleModelSpaceHomeomorph_coe_symm : (TotalSpace.toProd H E).symm := rfl +theorem contMDiff_tangentBundleModelSpaceHomeomorph {n : ℕ∞} : + ContMDiff I.tangent (I.prod 𝓘(𝕜, E)) n + (tangentBundleModelSpaceHomeomorph I : TangentBundle I H → ModelProd H E) := by + apply contMDiff_iff.2 ⟨Homeomorph.continuous _, fun x y ↦ ?_⟩ + apply contDiffOn_id.congr + simp only [mfld_simps, mem_range, TotalSpace.toProd, Equiv.coe_fn_symm_mk, forall_exists_index, + Prod.forall, Prod.mk.injEq] + rintro a b x rfl + simp [PartialEquiv.prod] + +theorem contMDiff_tangentBundleModelSpaceHomeomorph_symm {n : ℕ∞} : + ContMDiff (I.prod 𝓘(𝕜, E)) I.tangent n + ((tangentBundleModelSpaceHomeomorph I).symm : ModelProd H E → TangentBundle I H) := by + apply contMDiff_iff.2 ⟨Homeomorph.continuous _, fun x y ↦ ?_⟩ + apply contDiffOn_id.congr + simp only [mfld_simps, mem_range, TotalSpace.toProd, Equiv.coe_fn_symm_mk, forall_exists_index, + Prod.forall, Prod.mk.injEq] + rintro a b x rfl + simp [PartialEquiv.prod] + exact ⟨rfl, rfl⟩ + +variable (H I) in +/-- In the tangent bundle to the model space, the second projection is smooth. -/ +lemma contMDiff_snd_tangentBundle_modelSpace {n : ℕ∞} : + ContMDiff I.tangent 𝓘(𝕜, E) n (fun (p : TangentBundle I H) ↦ p.2) := by + change ContMDiff I.tangent 𝓘(𝕜, E) n + ((id Prod.snd : ModelProd H E → E) ∘ (tangentBundleModelSpaceHomeomorph I)) + apply ContMDiff.comp (I' := I.prod 𝓘(𝕜, E)) + · convert contMDiff_snd + rw [chartedSpaceSelf_prod] + rfl + · exact contMDiff_tangentBundleModelSpaceHomeomorph + +/-- A vector field on a vector space is smooth in the manifold sense iff it is smoooth in the vector +space sense-/ +lemma contMDiffWithinAt_vectorSpace_iff_contDiffWithinAt + {V : Π (x : E), TangentSpace 𝓘(𝕜, E) x} {n : ℕ∞} {s : Set E} {x : E} : + ContMDiffWithinAt 𝓘(𝕜, E) 𝓘(𝕜, E).tangent n (fun x ↦ (V x : TangentBundle 𝓘(𝕜, E) E)) s x ↔ + ContDiffWithinAt 𝕜 n V s x := by + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · exact ContMDiffWithinAt.contDiffWithinAt <| + (contMDiff_snd_tangentBundle_modelSpace E 𝓘(𝕜, E)).contMDiffAt.comp_contMDiffWithinAt _ h + · apply (Bundle.contMDiffWithinAt_totalSpace _).2 + refine ⟨contMDiffWithinAt_id, ?_⟩ + convert h.contMDiffWithinAt with y + simp + +/-- A vector field on a vector space is smooth in the manifold sense iff it is smoooth in the vector +space sense-/ +lemma contMDiffAt_vectorSpace_iff_contDiffAt + {V : Π (x : E), TangentSpace 𝓘(𝕜, E) x} {n : ℕ∞} {x : E} : + ContMDiffAt 𝓘(𝕜, E) 𝓘(𝕜, E).tangent n (fun x ↦ (V x : TangentBundle 𝓘(𝕜, E) E)) x ↔ + ContDiffAt 𝕜 n V x := by + simp only [← contMDiffWithinAt_univ, ← contDiffWithinAt_univ, + contMDiffWithinAt_vectorSpace_iff_contDiffWithinAt] + +/-- A vector field on a vector space is smooth in the manifold sense iff it is smoooth in the vector +space sense-/ +lemma contMDiffOn_vectorSpace_iff_contDiffOn + {V : Π (x : E), TangentSpace 𝓘(𝕜, E) x} {n : ℕ∞} {s : Set E} : + ContMDiffOn 𝓘(𝕜, E) 𝓘(𝕜, E).tangent n (fun x ↦ (V x : TangentBundle 𝓘(𝕜, E) E)) s ↔ + ContDiffOn 𝕜 n V s := by + simp only [ContMDiffOn, ContDiffOn, contMDiffWithinAt_vectorSpace_iff_contDiffWithinAt ] + +/-- A vector field on a vector space is smooth in the manifold sense iff it is smoooth in the vector +space sense-/ +lemma contMDiff_vectorSpace_iff_contDiff + {V : Π (x : E), TangentSpace 𝓘(𝕜, E) x} {n : ℕ∞} : + ContMDiff 𝓘(𝕜, E) 𝓘(𝕜, E).tangent n (fun x ↦ (V x : TangentBundle 𝓘(𝕜, E) E)) ↔ + ContDiff 𝕜 n V := by + simp only [← contMDiffOn_univ, ← contDiffOn_univ, contMDiffOn_vectorSpace_iff_contDiffOn] + section inTangentCoordinates variable {N : Type*} -/-- The map `in_coordinates` for the tangent bundle is trivial on the model spaces -/ +/-- The map `inCoordinates` for the tangent bundle is trivial on the model spaces -/ theorem inCoordinates_tangent_bundle_core_model_space (x₀ x : H) (y₀ y : H') (ϕ : E →L[𝕜] E') : inCoordinates E (TangentSpace I) E' (TangentSpace I') x₀ x y₀ y ϕ = ϕ := by erw [VectorBundleCore.inCoordinates_eq] <;> try trivial @@ -397,6 +477,10 @@ theorem inTangentCoordinates_model_space (f : N → H) (g : N → H') (ϕ : N simp (config := { unfoldPartialApp := true }) only [inTangentCoordinates, inCoordinates_tangent_bundle_core_model_space] +/-- To write a linear map between tangent spaces in coordinates amounts to precomposing and +postcomposing it with suitable coordinate changes. For a concrete version expressing the +change of coordinates as derivatives of extended charts, +see `inTangentCoordinates_eq_mfderiv_comp`. -/ theorem inTangentCoordinates_eq (f : N → M) (g : N → M') (ϕ : N → E →L[𝕜] E') {x₀ x : N} (hx : f x ∈ (chartAt H (f x₀)).source) (hy : g x ∈ (chartAt H' (g x₀)).source) : inTangentCoordinates I I' f g ϕ x₀ x = From e0fb3b122bf4cc764c699d9b980730a44226a533 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Mon, 9 Dec 2024 18:26:28 +0000 Subject: [PATCH 601/829] feat(RingTheory/Ideal/Maps): generalizations from Ring to Semiring (#19715) The results generalized are about bijective ring homs. Also add lemmas `comap.isMaximal` and `IsMaximal.comap_piEvalRingHom`, which says that the preimage of a maximal ideal under the projection from a product semiring is also maximal. --- Mathlib/RingTheory/Ideal/Maps.lean | 109 ++++++++++++++---------- Mathlib/RingTheory/Jacobson/Ring.lean | 2 +- Mathlib/RingTheory/Nullstellensatz.lean | 19 ++--- 3 files changed, 74 insertions(+), 56 deletions(-) diff --git a/Mathlib/RingTheory/Ideal/Maps.lean b/Mathlib/RingTheory/Ideal/Maps.lean index e2a631a34c0a9..8c11e7e809c5d 100644 --- a/Mathlib/RingTheory/Ideal/Maps.lean +++ b/Mathlib/RingTheory/Ideal/Maps.lean @@ -321,6 +321,19 @@ theorem map_eq_submodule_map (f : R →+* S) [h : RingHomSurjective f] (I : Idea I.map f = Submodule.map f.toSemilinearMap I := Submodule.ext fun _ => mem_map_iff_of_surjective f h.1 +open Function in +theorem IsMaximal.comap_piEvalRingHom {ι : Type*} {R : ι → Type*} [∀ i, Semiring (R i)] + {i : ι} {I : Ideal (R i)} (h : I.IsMaximal) : (I.comap <| Pi.evalRingHom R i).IsMaximal := by + refine isMaximal_iff.mpr ⟨I.ne_top_iff_one.mp h.ne_top, fun J x le hxI hxJ ↦ ?_⟩ + have ⟨r, y, hy, eq⟩ := h.exists_inv hxI + classical + convert J.add_mem (J.mul_mem_left (update 0 i r) hxJ) + (b := update 1 i y) (le <| by apply update_same i y 1 ▸ hy) + ext j + obtain rfl | ne := eq_or_ne j i + · simpa [eq_comm] using eq + · simp [update_noteq ne] + end Surjective section Injective @@ -353,7 +366,7 @@ theorem comap_of_equiv {I : Ideal R} (f : R ≃+* S) : /-- If `f : R ≃+* S` is a ring isomorphism and `I : Ideal R`, then `map f I = comap f.symm I`. -/ theorem map_comap_of_equiv {I : Ideal R} (f : R ≃+* S) : I.map (f : R →+* S) = I.comap f.symm := le_antisymm (Ideal.map_le_comap_of_inverse _ _ _ (Equiv.left_inv' _)) - (Ideal.comap_le_map_of_inverse _ _ _ (Equiv.right_inv' _)) + (Ideal.comap_le_map_of_inverse _ _ _ (Equiv.right_inv' _)) /-- If `f : R ≃+* S` is a ring isomorphism and `I : Ideal R`, then `comap f.symm I = map f I`. -/ @[simp] @@ -385,6 +398,56 @@ theorem mem_map_of_equiv {E : Type*} [EquivLike E R S] [RingEquivClass E R S] (e · rintro ⟨x, hx, rfl⟩ exact mem_map_of_mem e hx +section Bijective + +variable (hf : Function.Bijective f) {I : Ideal R} {K : Ideal S} +include hf + +/-- Special case of the correspondence theorem for isomorphic rings -/ +def relIsoOfBijective : Ideal S ≃o Ideal R where + toFun := comap f + invFun := map f + left_inv := map_comap_of_surjective _ hf.2 + right_inv J := + le_antisymm + (fun _ h ↦ have ⟨y, hy, eq⟩ := (mem_map_iff_of_surjective _ hf.2).mp h; hf.1 eq ▸ hy) + le_comap_map + map_rel_iff' {_ _} := by + refine ⟨fun h ↦ ?_, comap_mono⟩ + have := map_mono (f := f) h + simpa only [Equiv.coe_fn_mk, map_comap_of_surjective f hf.2] using this + +theorem comap_le_iff_le_map : comap f K ≤ I ↔ K ≤ map f I := + ⟨fun h => le_map_of_comap_le_of_surjective f hf.right h, fun h => + (relIsoOfBijective f hf).right_inv I ▸ comap_mono h⟩ + +lemma comap_map_of_bijective : (I.map f).comap f = I := + le_antisymm ((comap_le_iff_le_map f hf).mpr fun _ ↦ id) le_comap_map + +theorem isMaximal_map_iff_of_bijective : IsMaximal (map f I) ↔ IsMaximal I := by + simpa only [isMaximal_def] using (relIsoOfBijective _ hf).symm.isCoatom_iff _ + +theorem isMaximal_comap_iff_of_bijective : IsMaximal (comap f K) ↔ IsMaximal K := by + simpa only [isMaximal_def] using (relIsoOfBijective _ hf).isCoatom_iff _ + +alias ⟨_, IsMaximal.map_bijective⟩ := isMaximal_map_iff_of_bijective +alias ⟨_, IsMaximal.comap_bijective⟩ := isMaximal_comap_iff_of_bijective + +/-- A ring isomorphism sends a maximal ideal to a maximal ideal. -/ +instance map_isMaximal_of_equiv {E : Type*} [EquivLike E R S] [RingEquivClass E R S] (e : E) + {p : Ideal R} [hp : p.IsMaximal] : (map e p).IsMaximal := + hp.map_bijective e (EquivLike.bijective e) + +theorem isMaximal_iff_of_bijective : (⊥ : Ideal R).IsMaximal ↔ (⊥ : Ideal S).IsMaximal := + ⟨fun h ↦ map_bot (f := f) ▸ h.map_bijective f hf, fun h ↦ have e := RingEquiv.ofBijective f hf + map_bot (f := e.symm) ▸ h.map_bijective _ e.symm.bijective⟩ + +@[deprecated (since := "2024-12-07")] alias map.isMaximal := IsMaximal.map_bijective +@[deprecated (since := "2024-12-07")] alias comap.isMaximal := IsMaximal.comap_bijective +@[deprecated (since := "2024-12-07")] alias RingEquiv.bot_maximal_iff := isMaximal_iff_of_bijective + +end Bijective + end Semiring section Ring @@ -431,7 +494,7 @@ theorem map_eq_top_or_isMaximal_of_surjective (hf : Function.Surjective f) {I : · exact map_le_iff_le_comap.1 (le_of_lt hJ) · exact fun h => hJ.right (le_map_of_comap_le_of_surjective f hf (le_of_eq h.symm)) -theorem comap_isMaximal_of_surjective (hf : Function.Surjective f) {K : Ideal S} [H : IsMaximal K]: +theorem comap_isMaximal_of_surjective (hf : Function.Surjective f) {K : Ideal S} [H : IsMaximal K] : IsMaximal (comap f K) := by refine ⟨⟨comap_ne_top _ H.1.1, fun J hJ => ?_⟩⟩ suffices map f J = ⊤ by @@ -459,48 +522,6 @@ theorem comap_le_comap_iff_of_surjective (hf : Function.Surjective f) (I J : Ide end Surjective -section Bijective - -/-- Special case of the correspondence theorem for isomorphic rings -/ -def relIsoOfBijective (hf : Function.Bijective f) : Ideal S ≃o Ideal R where - toFun := comap f - invFun := map f - left_inv := (relIsoOfSurjective f hf.right).left_inv - right_inv J := - Subtype.ext_iff.1 - ((relIsoOfSurjective f hf.right).right_inv ⟨J, comap_bot_le_of_injective f hf.left⟩) - map_rel_iff' {_ _} := (relIsoOfSurjective f hf.right).map_rel_iff' - -theorem comap_le_iff_le_map (hf : Function.Bijective f) {I : Ideal R} {K : Ideal S} : - comap f K ≤ I ↔ K ≤ map f I := - ⟨fun h => le_map_of_comap_le_of_surjective f hf.right h, fun h => - (relIsoOfBijective f hf).right_inv I ▸ comap_mono h⟩ - -lemma comap_map_of_bijective (hf : Function.Bijective f) {I : Ideal R} : - (I.map f).comap f = I := - le_antisymm ((comap_le_iff_le_map f hf).mpr fun _ ↦ id) le_comap_map - -theorem map.isMaximal (hf : Function.Bijective f) {I : Ideal R} (H : IsMaximal I) : - IsMaximal (map f I) := by - refine - or_iff_not_imp_left.1 (map_eq_top_or_isMaximal_of_surjective f hf.right H) fun h => H.1.1 ?_ - calc - I = comap f (map f I) := ((relIsoOfBijective f hf).right_inv I).symm - _ = comap f ⊤ := by rw [h] - _ = ⊤ := by rw [comap_top] - -/-- A ring isomorphism sends a maximal ideal to a maximal ideal. -/ -instance map_isMaximal_of_equiv {E : Type*} [EquivLike E R S] [RingEquivClass E R S] (e : E) - {p : Ideal R} [hp : p.IsMaximal] : (map e p).IsMaximal := - map.isMaximal e (EquivLike.bijective e) hp - -end Bijective - -theorem RingEquiv.bot_maximal_iff (e : R ≃+* S) : - (⊥ : Ideal R).IsMaximal ↔ (⊥ : Ideal S).IsMaximal := - ⟨fun h => map_bot (f := e.toRingHom) ▸ map.isMaximal e.toRingHom e.bijective h, fun h => - map_bot (f := e.symm.toRingHom) ▸ map.isMaximal e.symm.toRingHom e.symm.bijective h⟩ - end Ring section CommRing diff --git a/Mathlib/RingTheory/Jacobson/Ring.lean b/Mathlib/RingTheory/Jacobson/Ring.lean index a8165786b21e0..958012570ae25 100644 --- a/Mathlib/RingTheory/Jacobson/Ring.lean +++ b/Mathlib/RingTheory/Jacobson/Ring.lean @@ -492,7 +492,7 @@ theorem isMaximal_comap_C_of_isMaximal [IsJacobsonRing R] [Nontrivial R] rw [(map_bot.symm : (⊥ : Ideal (Localization M')) = Ideal.map (algebraMap (R[X] ⧸ P) (Localization M')) ⊥)] let bot_maximal := (bot_quotient_isMaximal_iff _).mpr hP - refine map.isMaximal (algebraMap (R[X] ⧸ P) (Localization M')) ?_ bot_maximal + refine bot_maximal.map_bijective (algebraMap (R[X] ⧸ P) (Localization M')) ?_ apply IsField.localization_map_bijective hM' rwa [← Quotient.maximal_ideal_iff_isField_quotient, ← bot_quotient_isMaximal_iff] diff --git a/Mathlib/RingTheory/Nullstellensatz.lean b/Mathlib/RingTheory/Nullstellensatz.lean index 6e41ad765593b..a0a41f1573af9 100644 --- a/Mathlib/RingTheory/Nullstellensatz.lean +++ b/Mathlib/RingTheory/Nullstellensatz.lean @@ -96,17 +96,14 @@ theorem mem_vanishingIdeal_singleton_iff (x : σ → k) (p : MvPolynomial σ k) instance vanishingIdeal_singleton_isMaximal {x : σ → k} : (vanishingIdeal {x} : Ideal (MvPolynomial σ k)).IsMaximal := by - have : MvPolynomial σ k ⧸ vanishingIdeal {x} ≃+* k := - RingEquiv.ofBijective - (Ideal.Quotient.lift _ (eval x) fun p h => (mem_vanishingIdeal_singleton_iff x p).mp h) - (by - refine - ⟨(injective_iff_map_eq_zero _).mpr fun p hp => ?_, fun z => - ⟨(Ideal.Quotient.mk (vanishingIdeal {x} : Ideal (MvPolynomial σ k))) (C z), by simp⟩⟩ - obtain ⟨q, rfl⟩ := Ideal.Quotient.mk_surjective p - rwa [Ideal.Quotient.lift_mk, ← mem_vanishingIdeal_singleton_iff, - ← Quotient.eq_zero_iff_mem] at hp) - rw [← bot_quotient_isMaximal_iff, RingEquiv.bot_maximal_iff this] + have : Function.Bijective + (Ideal.Quotient.lift _ (eval x) fun p h ↦ (mem_vanishingIdeal_singleton_iff x p).mp h) := by + refine ⟨(injective_iff_map_eq_zero _).mpr fun p hp ↦ ?_, fun z ↦ + ⟨(Ideal.Quotient.mk (vanishingIdeal {x} : Ideal (MvPolynomial σ k))) (C z), by simp⟩⟩ + obtain ⟨q, rfl⟩ := Ideal.Quotient.mk_surjective p + rwa [Ideal.Quotient.lift_mk, ← mem_vanishingIdeal_singleton_iff, + ← Quotient.eq_zero_iff_mem] at hp + rw [← bot_quotient_isMaximal_iff, isMaximal_iff_of_bijective _ this] exact bot_isMaximal theorem radical_le_vanishingIdeal_zeroLocus (I : Ideal (MvPolynomial σ k)) : From cd316de5d3fbd8a9ccc91b61e421c1bc2d58f438 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Mon, 9 Dec 2024 19:13:48 +0000 Subject: [PATCH 602/829] =?UTF-8?q?chore:=20reduce=20type=20class=20assump?= =?UTF-8?q?tions=20for=20`rpow`=20lemmas=20in=20C=E2=8B=86-algebras=20(#19?= =?UTF-8?q?843)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- .../SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean index 43d2501f56550..487d14cd2871b 100644 --- a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean +++ b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean @@ -66,7 +66,7 @@ namespace CFC section NonUnital -variable {A : Type*} [PartialOrder A] [NonUnitalNormedRing A] [StarRing A] +variable {A : Type*} [PartialOrder A] [NonUnitalRing A] [TopologicalSpace A] [StarRing A] [Module ℝ≥0 A] [SMulCommClass ℝ≥0 A A] [IsScalarTower ℝ≥0 A A] [NonUnitalContinuousFunctionalCalculus ℝ≥0 (fun (a : A) => 0 ≤ a)] @@ -210,8 +210,8 @@ end NonUnital section Unital -variable {A : Type*} [PartialOrder A] [NormedRing A] [StarRing A] - [NormedAlgebra ℝ A] [ContinuousFunctionalCalculus ℝ≥0 (fun (a : A) => 0 ≤ a)] +variable {A : Type*} [PartialOrder A] [Ring A] [StarRing A] [TopologicalSpace A] + [Algebra ℝ A] [ContinuousFunctionalCalculus ℝ≥0 (fun (a : A) => 0 ≤ a)] /- ## `rpow` -/ From d1e3519e88429fe05253ed3fa52c16ddc3f2377a Mon Sep 17 00:00:00 2001 From: grunweg Date: Mon, 9 Dec 2024 19:23:39 +0000 Subject: [PATCH 603/829] perf: mark some definitions as non-computable (#19740) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit See individual commits for details about each change. This saves about 300*10⁹ instructions on all of mathlib. --- Mathlib/Algebra/Category/Ring/Under/Limits.lean | 8 ++++++-- Mathlib/AlgebraicTopology/DoldKan/GammaCompN.lean | 2 ++ Mathlib/CategoryTheory/Comma/Over.lean | 6 ++++-- Mathlib/CategoryTheory/Limits/Opposites.lean | 8 ++++++-- Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean | 2 +- 5 files changed, 19 insertions(+), 7 deletions(-) diff --git a/Mathlib/Algebra/Category/Ring/Under/Limits.lean b/Mathlib/Algebra/Category/Ring/Under/Limits.lean index 29babdcccee80..24c129198c8b4 100644 --- a/Mathlib/Algebra/Category/Ring/Under/Limits.lean +++ b/Mathlib/Algebra/Category/Ring/Under/Limits.lean @@ -82,6 +82,7 @@ def tensorProductFanIsLimit [Finite ι] : IsLimit (tensorProductFan S P) := (IsLimit.equivIsoLimit (tensorProductFanIso P)).symm (Under.piFanIsLimit _) /-- `tensorProd R S` preserves the limit of the canonical fan on `P`. -/ +noncomputable -- marked noncomputable for performance (only) def piFanTensorProductIsLimit [Finite ι] : IsLimit ((tensorProd R S).mapCone (Under.piFan P)) := (isLimitMapConeFanMkEquiv (tensorProd R S) P _).symm <| tensorProductFanIsLimit P @@ -135,7 +136,8 @@ lemma equalizerFork'_ι {A B : Type u} [CommRing A] [CommRing B] [Algebra R A] [ (Under.equalizerFork' f g).ι = (AlgHom.equalizer f g).val.toUnder := rfl /-- The canonical fork on `f g : A ⟶ B` is limiting. -/ -def equalizerForkIsLimit {A B : Under R} (f g : A ⟶ B) : +-- marked noncomputable for performance (only) +noncomputable def equalizerForkIsLimit {A B : Under R} (f g : A ⟶ B) : IsLimit (Under.equalizerFork f g) := isLimitOfReflects (Under.forget R) <| (isLimitMapConeForkEquiv (Under.forget R) (equalizer_comp f g)).invFun <| @@ -161,7 +163,8 @@ lemma tensorProdEqualizer_ι {A B : Under R} (f g : A ⟶ B) : rfl /-- If `S` is `R`-flat, `S ⊗[R] eq(f, g)` is isomorphic to `eq(𝟙 ⊗[R] f, 𝟙 ⊗[R] g)`. -/ -def equalizerForkTensorProdIso [Module.Flat R S] {A B : Under R} (f g : A ⟶ B) : +-- marked noncomputable for performance (only) +noncomputable def equalizerForkTensorProdIso [Module.Flat R S] {A B : Under R} (f g : A ⟶ B) : tensorProdEqualizer f g ≅ Under.equalizerFork' (Algebra.TensorProduct.map (AlgHom.id S S) (toAlgHom f)) (Algebra.TensorProduct.map (AlgHom.id S S) (toAlgHom g)) := @@ -170,6 +173,7 @@ def equalizerForkTensorProdIso [Module.Flat R S] {A B : Under R} (f g : A ⟶ B) apply AlgHom.coe_tensorEqualizer /-- If `S` is `R`-flat, `tensorProd R S` preserves the equalizer of `f` and `g`. -/ +noncomputable -- marked noncomputable for performance (only) def tensorProdMapEqualizerForkIsLimit [Module.Flat R S] {A B : Under R} (f g : A ⟶ B) : IsLimit ((tensorProd R S).mapCone <| Under.equalizerFork f g) := (isLimitMapConeForkEquiv (tensorProd R S) _).symm <| diff --git a/Mathlib/AlgebraicTopology/DoldKan/GammaCompN.lean b/Mathlib/AlgebraicTopology/DoldKan/GammaCompN.lean index b6025650b82bb..3081d5036b92a 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/GammaCompN.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/GammaCompN.lean @@ -6,6 +6,7 @@ Authors: Joël Riou import Mathlib.AlgebraicTopology.DoldKan.FunctorGamma import Mathlib.AlgebraicTopology.DoldKan.SplitSimplicialObject import Mathlib.CategoryTheory.Idempotents.HomologicalComplex +import Mathlib.Tactic.SuppressCompilation /-! The counit isomorphism of the Dold-Kan equivalence @@ -17,6 +18,7 @@ and `N₂Γ₂ : Γ₂ ⋙ N₂ ≅ 𝟭 (Karoubi (ChainComplex C ℕ))`. -/ +suppress_compilation noncomputable section diff --git a/Mathlib/CategoryTheory/Comma/Over.lean b/Mathlib/CategoryTheory/Comma/Over.lean index 9f9609c579450..8b852e7627a5b 100644 --- a/Mathlib/CategoryTheory/Comma/Over.lean +++ b/Mathlib/CategoryTheory/Comma/Over.lean @@ -852,7 +852,8 @@ def ofDiagEquivalence (X : T × T) : /-- A version of `StructuredArrow.ofDiagEquivalence` with the roles of the first and second projection swapped. -/ -def ofDiagEquivalence' (X : T × T) : +-- noncomputability is only for performance +noncomputable def ofDiagEquivalence' (X : T × T) : StructuredArrow X (Functor.diag _) ≌ StructuredArrow X.1 (Under.forget X.2) := (ofDiagEquivalence X).trans <| (ofStructuredArrowProjEquivalence (𝟭 T) X.1 X.2).trans <| @@ -919,7 +920,8 @@ def ofDiagEquivalence (X : T × T) : /-- A version of `CostructuredArrow.ofDiagEquivalence` with the roles of the first and second projection swapped. -/ -def ofDiagEquivalence' (X : T × T) : +-- noncomputability is only for performance +noncomputable def ofDiagEquivalence' (X : T × T) : CostructuredArrow (Functor.diag _) X ≌ CostructuredArrow (Over.forget X.2) X.1 := (ofDiagEquivalence X).trans <| (ofCostructuredArrowProjEquivalence (𝟭 T) X.1 X.2).trans <| diff --git a/Mathlib/CategoryTheory/Limits/Opposites.lean b/Mathlib/CategoryTheory/Limits/Opposites.lean index 493dec55ff8dc..eeffcd5c0cf60 100644 --- a/Mathlib/CategoryTheory/Limits/Opposites.lean +++ b/Mathlib/CategoryTheory/Limits/Opposites.lean @@ -570,7 +570,8 @@ instance : HasProduct (op <| Z ·) := hasLimitOfIso def Cofan.op (c : Cofan Z) : Fan (op <| Z ·) := Fan.mk _ (fun a ↦ (c.inj a).op) /-- If a `Cofan` is colimit, then its opposite is limit. -/ -def Cofan.IsColimit.op {c : Cofan Z} (hc : IsColimit c) : IsLimit c.op := by +-- noncomputability is just for performance (compilation takes a while) +noncomputable def Cofan.IsColimit.op {c : Cofan Z} (hc : IsColimit c) : IsLimit c.op := by let e : Discrete.functor (Opposite.op <| Z ·) ≅ (Discrete.opposite α).inverse ⋙ (Discrete.functor Z).op := Discrete.natIso (fun _ ↦ Iso.refl _) refine IsLimit.ofIsoLimit ((IsLimit.postcomposeInvEquiv e _).2 @@ -666,7 +667,8 @@ instance : HasCoproduct (op <| Z ·) := hasColimitOfIso def Fan.op (f : Fan Z) : Cofan (op <| Z ·) := Cofan.mk _ (fun a ↦ (f.proj a).op) /-- If a `Fan` is limit, then its opposite is colimit. -/ -def Fan.IsLimit.op {f : Fan Z} (hf : IsLimit f) : IsColimit f.op := by +-- noncomputability is just for performance (compilation takes a while) +noncomputable def Fan.IsLimit.op {f : Fan Z} (hf : IsLimit f) : IsColimit f.op := by let e : Discrete.functor (Opposite.op <| Z ·) ≅ (Discrete.opposite α).inverse ⋙ (Discrete.functor Z).op := Discrete.natIso (fun _ ↦ Iso.refl _) refine IsColimit.ofIsoColimit ((IsColimit.precomposeHomEquiv e _).2 @@ -950,6 +952,7 @@ def unopOp {X Y Z : Cᵒᵖ} {f : X ⟶ Y} {g : X ⟶ Z} (c : PushoutCocone f g) /-- A pushout cone is a colimit cocone if and only if the corresponding pullback cone in the opposite category is a limit cone. -/ +noncomputable -- just for performance; compilation takes several seconds def isColimitEquivIsLimitOp {X Y Z : C} {f : X ⟶ Y} {g : X ⟶ Z} (c : PushoutCocone f g) : IsColimit c ≃ IsLimit c.op := by apply equivOfSubsingletonOfSubsingleton @@ -963,6 +966,7 @@ def isColimitEquivIsLimitOp {X Y Z : C} {f : X ⟶ Y} {g : X ⟶ Z} (c : Pushout /-- A pushout cone is a colimit cocone in `Cᵒᵖ` if and only if the corresponding pullback cone in `C` is a limit cone. -/ +noncomputable -- just for performance; compilation takes several seconds def isColimitEquivIsLimitUnop {X Y Z : Cᵒᵖ} {f : X ⟶ Y} {g : X ⟶ Z} (c : PushoutCocone f g) : IsColimit c ≃ IsLimit c.unop := by apply equivOfSubsingletonOfSubsingleton diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean index dfe5ff42f7453..29f15f3221a2d 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean @@ -288,7 +288,7 @@ def whiskerToCocone {f : J → C} (c : Bicone f) (g : K ≃ J) : Cocones.ext (Iso.refl _) (by aesop_cat) /-- Whiskering a bicone with an equivalence between types preserves being a bilimit bicone. -/ -def whiskerIsBilimitIff {f : J → C} (c : Bicone f) (g : K ≃ J) : +noncomputable def whiskerIsBilimitIff {f : J → C} (c : Bicone f) (g : K ≃ J) : (c.whisker g).IsBilimit ≃ c.IsBilimit := by refine equivOfSubsingletonOfSubsingleton (fun hc => ⟨?_, ?_⟩) fun hc => ⟨?_, ?_⟩ · let this := IsLimit.ofIsoLimit hc.isLimit (Bicone.whiskerToCone c g) From fb9a612d26b507a6a183124d6b87e3996c30e33c Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Mon, 9 Dec 2024 19:23:41 +0000 Subject: [PATCH 604/829] chore: simplify variables in LinearAlgebra.Dimension.Torsion.Finite (#19839) --- .../Dimension/Torsion/Finite.lean | 21 +++---------------- 1 file changed, 3 insertions(+), 18 deletions(-) diff --git a/Mathlib/LinearAlgebra/Dimension/Torsion/Finite.lean b/Mathlib/LinearAlgebra/Dimension/Torsion/Finite.lean index 6929e94e571f5..3271bdb4aa7c9 100644 --- a/Mathlib/LinearAlgebra/Dimension/Torsion/Finite.lean +++ b/Mathlib/LinearAlgebra/Dimension/Torsion/Finite.lean @@ -11,29 +11,14 @@ import Mathlib.LinearAlgebra.Dimension.Finite -/ -universe u v +variable {R M : Type*} [CommRing R] [IsDomain R] [AddCommGroup M] [Module R M] -variable {R : Type u} {M : Type v} -variable [Ring R] [AddCommGroup M] - -section RankZero - -variable [Nontrivial R] - -lemma rank_eq_zero_iff_isTorsion {R M} [CommRing R] [IsDomain R] [AddCommGroup M] [Module R M] : - Module.rank R M = 0 ↔ Module.IsTorsion R M := by +lemma rank_eq_zero_iff_isTorsion: Module.rank R M = 0 ↔ Module.IsTorsion R M := by rw [Module.IsTorsion, rank_eq_zero_iff] simp [mem_nonZeroDivisors_iff_ne_zero] -end RankZero - -section StrongRankCondition - /-- The `StrongRankCondition` is automatic. See `commRing_strongRankCondition`. -/ -theorem Module.finrank_eq_zero_iff_isTorsion {R} [CommRing R] [StrongRankCondition R] - [IsDomain R] [Module R M] [Module.Finite R M] : +theorem Module.finrank_eq_zero_iff_isTorsion [StrongRankCondition R] [Module.Finite R M] : finrank R M = 0 ↔ Module.IsTorsion R M := by rw [← rank_eq_zero_iff_isTorsion (R := R), ← finrank_eq_rank] norm_cast - -end StrongRankCondition From d0592cc111886eed7510d47018a0cdca5f28b38a Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Tue, 10 Dec 2024 01:28:30 +0000 Subject: [PATCH 605/829] chore(.github/workflows/discover-lean-pr-testing.yml): change zulip message topic (#19822) --- .github/workflows/discover-lean-pr-testing.yml | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/.github/workflows/discover-lean-pr-testing.yml b/.github/workflows/discover-lean-pr-testing.yml index fa3dcad9ae5fe..04b70247d680f 100644 --- a/.github/workflows/discover-lean-pr-testing.yml +++ b/.github/workflows/discover-lean-pr-testing.yml @@ -133,8 +133,10 @@ jobs: organization-url: 'https://leanprover.zulipchat.com' to: 'nightly-testing' type: 'stream' - topic: 'Mathlib status updates' + topic: 'Mergeable lean testing PRs' content: | We will need to merge the following branches into `nightly-testing`: ${{ steps.find-branches.outputs.branches }} + + This can be done using the script `scripts/merge-lean-testing-pr.sh`, which takes 1 lean-PR number as argument. From 721b993578ff801d19aa0e4dbcf30c9fe0730432 Mon Sep 17 00:00:00 2001 From: grunweg Date: Tue, 10 Dec 2024 01:37:15 +0000 Subject: [PATCH 606/829] perf(Linter/HaveLetLinter): some local optimisations (#19730) Suggested by `kim-em` on [Zulip](https://leanprover.zulipchat.com/#narrow/channel/144837-PR-reviews/topic/.2312190.20.60have.60.20vs.20.60let.60.20linter/near/486013460). --- Mathlib/Tactic/Linter/HaveLetLinter.lean | 52 +++++++++++------------- 1 file changed, 24 insertions(+), 28 deletions(-) diff --git a/Mathlib/Tactic/Linter/HaveLetLinter.lean b/Mathlib/Tactic/Linter/HaveLetLinter.lean index fe586741660b6..b88a5344a4265 100644 --- a/Mathlib/Tactic/Linter/HaveLetLinter.lean +++ b/Mathlib/Tactic/Linter/HaveLetLinter.lean @@ -66,21 +66,21 @@ def InfoTree.foldInfoM {α m} [Monad m] (f : ContextInfo → Info → α → m InfoTree → m α := InfoTree.foldInfo (fun ctx i ma => do f ctx i (← ma)) (pure init) -/-- given a `ContextInfo`, a `LocalContext` and an `Array` of `Expr`essions `es`, -`areProp_toFormat` creates a `MetaM` context, and returns an array of pairs consisting of -* a `Bool`ean, answering the question of whether the Type of `e` is a `Prop` or not, and -* the pretty-printed `Format` of `e` -for each `Expr`ession `e` in `es`. -Concretely, `areProp_toFormat` runs `inferType` in `CommandElabM`. +/-- given a `ContextInfo`, a `LocalContext` and an `Array` of `Expr`essions `es` with a `Name`, +`toFormat_propTypes` creates a `MetaM` context, and returns an array of +the pretty-printed `Format` of `e`, together with the (unchanged) name +for each `Expr`ession `e` in `es` whose type is a `Prop`. + +Concretely, `toFormat_propTypes` runs `inferType` in `CommandElabM`. This is the kind of monadic lift that `nonPropHaves` uses to decide whether the Type of a `have` is in `Prop` or not. The output `Format` is just so that the linter displays a better message. -/ -def areProp_toFormat (ctx : ContextInfo) (lc : LocalContext) (es : Array Expr) : - CommandElabM (Array (Bool × Format)) := do +def toFormat_propTypes (ctx : ContextInfo) (lc : LocalContext) (es : Array (Expr × Name)) : + CommandElabM (Array (Format × Name)) := do ctx.runMetaM lc do - es.mapM fun e => do + es.filterMapM fun (e, name) ↦ do let typ ← inferType (← instantiateMVars e) - return (typ.isProp, ← ppExpr e) + if typ.isProp then return none else return (← ppExpr e, name) /-- returns the `have` syntax whose corresponding hypothesis does not have Type `Prop` and also a `Format`ted version of the corresponding Type. -/ @@ -92,26 +92,22 @@ def nonPropHaves : InfoTree → CommandElabM (Array (Syntax × Format)) := let .original .. := stx.getHeadInfo | return #[] unless isHave? stx do return #[] let mctx := i.mctxAfter - let mvdecls := (i.goalsAfter.map (mctx.decls.find? ·)).reduceOption - -- we extract the `MetavarDecl` with largest index after a `have`, since this one - -- holds information about the metavariable where `have` introduces the new hypothesis. - let largestIdx := mvdecls.toArray.qsort (·.index > ·.index) - -- the relevant `LocalContext` - let lc := (largestIdx.getD 0 default).lctx + let mvdecls := i.goalsAfter.filterMap (mctx.decls.find? ·) + -- We extract the `MetavarDecl` with largest index after a `have`, since this one + -- holds information about the metavariable where `have` introduces the new hypothesis, + -- and determine the relevant `LocalContext`. + let lc := mvdecls.toArray.getMax? (·.index < ·.index) |>.getD default |>.lctx -- we also accumulate all `fvarId`s from all local contexts before the use of `have` -- so that we can then isolate the `fvarId`s that are created by `have` - let oldMvdecls := (i.goalsBefore.map (mctx.decls.find? ·)).reduceOption - let oldLctx := oldMvdecls.map (·.lctx) - let oldDecls := (oldLctx.map (·.decls.toList.reduceOption)).flatten - let oldFVars := oldDecls.map (·.fvarId) - -- `newDecls` are the local declarations whose `FVarID` did not exist before the `have` - -- effectively they are the declarations that we want to test for being in `Prop` or not. + let oldMvdecls := i.goalsBefore.filterMap (mctx.decls.find? ·) + let oldFVars := (oldMvdecls.map (·.lctx.decls.toList.reduceOption)).flatten.map (·.fvarId) + -- `newDecls` are the local declarations whose `FVarID` did not exist before the `have`. + -- Effectively they are the declarations that we want to test for being in `Prop` or not. let newDecls := lc.decls.toList.reduceOption.filter (! oldFVars.contains ·.fvarId) - -- now, we get the `MetaM` state up and running to find the types of each entry of `newDecls` - let fmts ← areProp_toFormat ctx lc (newDecls.map (·.type)).toArray - let (_propFmts, typeFmts) := (fmts.zip (newDecls.map (·.userName)).toArray).partition (·.1.1) - -- everything that is a Type triggers a warning on `have` - return typeFmts.map fun ((_, fmt), na) => (stx, f!"{na} : {fmt}")) + -- Now, we get the `MetaM` state up and running to find the types of each entry of `newDecls`. + -- For each entry which is a `Type`, we print a warning on `have`. + let fmts ← toFormat_propTypes ctx lc (newDecls.map (fun e ↦ (e.type, e.userName))).toArray + return fmts.map fun (fmt, na) ↦ (stx, f!"{na} : {fmt}")) /-- The main implementation of the `have` vs `let` linter. -/ def haveLetLinter : Linter where run := withSetOptionIn fun _stx => do @@ -120,7 +116,7 @@ def haveLetLinter : Linter where run := withSetOptionIn fun _stx => do return unless gh == 1 && (← MonadState.get).messages.unreported.isEmpty do let trees ← getInfoTrees - for t in trees.toArray do + for t in trees do for (s, fmt) in ← nonPropHaves t do -- Since the linter option is not in `Bool`, the standard `Linter.logLint` does not work. -- We emulate it with `logWarningAt` From 6220f584f70ec93e08a33d4efb9e9015517dc605 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 10 Dec 2024 02:17:59 +0000 Subject: [PATCH 607/829] chore(Util/Superscript): use `register_parser_alias` (#19009) This should mean exactly the same thing, with some automatic name resolution. --- Mathlib/Util/Superscript.lean | 8 ++------ 1 file changed, 2 insertions(+), 6 deletions(-) diff --git a/Mathlib/Util/Superscript.lean b/Mathlib/Util/Superscript.lean index 85d3b2fecae7f..e322df63fe38f 100644 --- a/Mathlib/Util/Superscript.lean +++ b/Mathlib/Util/Superscript.lean @@ -277,11 +277,7 @@ def subscript.parenthesizer := Superscript.scriptParser.parenthesizer ``subscrip def subscript.formatter := Superscript.scriptParser.formatter "subscript" .subscript ``subscript initialize - registerAlias `superscript ``superscript superscript - registerAliasCore Formatter.formatterAliasesRef `superscript superscript.formatter - registerAliasCore Parenthesizer.parenthesizerAliasesRef `superscript superscript.parenthesizer - registerAlias `subscript ``subscript subscript - registerAliasCore Formatter.formatterAliasesRef `subscript subscript.formatter - registerAliasCore Parenthesizer.parenthesizerAliasesRef `subscript subscript.parenthesizer + register_parser_alias superscript + register_parser_alias subscript end Mathlib.Tactic From 8dd91ce473340d1d338c504c69dfa23ebd68ee34 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 10 Dec 2024 02:18:01 +0000 Subject: [PATCH 608/829] doc(SetTheory/ZFC/Ordinal): add docstrings to lemmas on transitive sets (#19266) --- Mathlib/SetTheory/ZFC/Ordinal.lean | 6 +++++- 1 file changed, 5 insertions(+), 1 deletion(-) diff --git a/Mathlib/SetTheory/ZFC/Ordinal.lean b/Mathlib/SetTheory/ZFC/Ordinal.lean index bd816b21e5dac..43a8b4f282040 100644 --- a/Mathlib/SetTheory/ZFC/Ordinal.lean +++ b/Mathlib/SetTheory/ZFC/Ordinal.lean @@ -34,7 +34,9 @@ variable {x y z : ZFSet.{u}} namespace ZFSet -/-- A transitive set is one where every element is a subset. -/ +/-- A transitive set is one where every element is a subset. + +This is equivalent to being an infinite-open interval in the transitive closure of membership. -/ def IsTransitive (x : ZFSet) : Prop := ∀ y ∈ x, y ⊆ x @@ -56,11 +58,13 @@ protected theorem IsTransitive.inter (hx : x.IsTransitive) (hy : y.IsTransitive) rw [mem_inter] at hz ⊢ exact ⟨hx.mem_trans hw hz.1, hy.mem_trans hw hz.2⟩ +/-- The union of a transitive set is transitive. -/ protected theorem IsTransitive.sUnion (h : x.IsTransitive) : (⋃₀ x : ZFSet).IsTransitive := fun y hy z hz => by rcases mem_sUnion.1 hy with ⟨w, hw, hw'⟩ exact mem_sUnion_of_mem hz (h.mem_trans hw' hw) +/-- The union of transitive sets is transitive. -/ theorem IsTransitive.sUnion' (H : ∀ y ∈ x, IsTransitive y) : (⋃₀ x : ZFSet).IsTransitive := fun y hy z hz => by rcases mem_sUnion.1 hy with ⟨w, hw, hw'⟩ From 69286d62bc85f8d57c9f402fa7aaeb3038bd840c Mon Sep 17 00:00:00 2001 From: grunweg Date: Tue, 10 Dec 2024 02:18:02 +0000 Subject: [PATCH 609/829] perf(Convex/Between): speed up two slow refines (#19742) Previously, each refine took more than 6s in CI. Extract the last goal as a `have` as use `apply` instead, which is near-instant. Note that `exact` or `refine` is still slow. Co-authored-by: grunweg --- Mathlib/Analysis/Convex/Between.lean | 10 ++++++---- 1 file changed, 6 insertions(+), 4 deletions(-) diff --git a/Mathlib/Analysis/Convex/Between.lean b/Mathlib/Analysis/Convex/Between.lean index 0359b1ad4b711..977702d9c6a6b 100644 --- a/Mathlib/Analysis/Convex/Between.lean +++ b/Mathlib/Analysis/Convex/Between.lean @@ -153,14 +153,16 @@ theorem Function.Injective.sbtw_map_iff {x y z : P} {f : P →ᵃ[R] P'} (hf : F @[simp] theorem AffineEquiv.wbtw_map_iff {x y z : P} (f : P ≃ᵃ[R] P') : Wbtw R (f x) (f y) (f z) ↔ Wbtw R x y z := by - refine Function.Injective.wbtw_map_iff (?_ : Function.Injective f.toAffineMap) - exact f.injective + have : Function.Injective f.toAffineMap := f.injective + -- `refine` or `exact` are very slow, `apply` is fast. Please check before golfing. + apply this.wbtw_map_iff @[simp] theorem AffineEquiv.sbtw_map_iff {x y z : P} (f : P ≃ᵃ[R] P') : Sbtw R (f x) (f y) (f z) ↔ Sbtw R x y z := by - refine Function.Injective.sbtw_map_iff (?_ : Function.Injective f.toAffineMap) - exact f.injective + have : Function.Injective f.toAffineMap := f.injective + -- `refine` or `exact` are very slow, `apply` is fast. Please check before golfing. + apply this.sbtw_map_iff @[simp] theorem wbtw_const_vadd_iff {x y z : P} (v : V) : From 61ce7cc64c47430f582de8b0397e5800768fd0ed Mon Sep 17 00:00:00 2001 From: Julian Berman Date: Tue, 10 Dec 2024 02:18:03 +0000 Subject: [PATCH 610/829] fix: update List.intervals docstring to point to new interval locations (#19794) (The old reference no longer exists.) --- Mathlib/Data/List/Intervals.lean | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/Mathlib/Data/List/Intervals.lean b/Mathlib/Data/List/Intervals.lean index 67e0ed0acd83e..bfc2d2a858d87 100644 --- a/Mathlib/Data/List/Intervals.lean +++ b/Mathlib/Data/List/Intervals.lean @@ -29,8 +29,9 @@ namespace List /-- `Ico n m` is the list of natural numbers `n ≤ x < m`. (Ico stands for "interval, closed-open".) -See also `Data/Set/Intervals.lean` for `Set.Ico`, modelling intervals in general preorders, and -`Multiset.Ico` and `Finset.Ico` for `n ≤ x < m` as a multiset or as a finset. +See also `Order/Interval/Basic.lean` for modelling intervals in general preorders, as well as +sibling definitions alongside it such as `Set.Ico`, `Multiset.Ico` and `Finset.Ico` +for sets, multisets and finite sets respectively. -/ def Ico (n m : ℕ) : List ℕ := range' n (m - n) From c4c2823fcb8b0eb245aeab75224a9554be0473ec Mon Sep 17 00:00:00 2001 From: Vasily Ilin Date: Tue, 10 Dec 2024 02:18:05 +0000 Subject: [PATCH 611/829] feat(Probability/Moments): add `mgf_smul_left` (#19798) Add `mgf_smul_left`. While there is a lemma [IndepFun.mgf_add](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Probability/Moments.html#ProbabilityTheory.IndepFun.mgf_add) that describes the additive behavior of the moment generating function, there is currently no lemma describing the (scalar-)multiplicative behavior. This lemma is one of the ~4 supporting lemmas we have proved on our way to formalizing the central limit theorem. We will make PRs for other lemmas as well. Co-authored-by: Siyuan Ge --- Mathlib/Probability/Moments.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/Probability/Moments.lean b/Mathlib/Probability/Moments.lean index 5cd7261ecfd42..1178661e1715f 100644 --- a/Mathlib/Probability/Moments.lean +++ b/Mathlib/Probability/Moments.lean @@ -184,6 +184,9 @@ theorem mgf_neg : mgf (-X) μ t = mgf X μ (-t) := by simp_rw [mgf, Pi.neg_apply theorem cgf_neg : cgf (-X) μ t = cgf X μ (-t) := by simp_rw [cgf, mgf_neg] +theorem mgf_smul_left (α : ℝ) : mgf (α • X) μ t = mgf X μ (α * t) := by + simp_rw [mgf, Pi.smul_apply, smul_eq_mul, mul_comm α t, mul_assoc] + /-- This is a trivial application of `IndepFun.comp` but it will come up frequently. -/ theorem IndepFun.exp_mul {X Y : Ω → ℝ} (h_indep : IndepFun X Y μ) (s t : ℝ) : IndepFun (fun ω => exp (s * X ω)) (fun ω => exp (t * Y ω)) μ := by From e96006be03a6675d2f7085115e7e7ee2ce4fa9fb Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 10 Dec 2024 02:18:06 +0000 Subject: [PATCH 612/829] chore(Order/SuccPred/CompleteLinearOrder): spacing fixes (#19799) --- .../Order/SuccPred/CompleteLinearOrder.lean | 56 ++++++++----------- 1 file changed, 24 insertions(+), 32 deletions(-) diff --git a/Mathlib/Order/SuccPred/CompleteLinearOrder.lean b/Mathlib/Order/SuccPred/CompleteLinearOrder.lean index eea7681180495..748703d2d7787 100644 --- a/Mathlib/Order/SuccPred/CompleteLinearOrder.lean +++ b/Mathlib/Order/SuccPred/CompleteLinearOrder.lean @@ -19,8 +19,7 @@ section ConditionallyCompleteLinearOrder variable [ConditionallyCompleteLinearOrder α] [Nonempty ι] {f : ι → α} {s : Set α} {x : α} lemma csSup_mem_of_not_isSuccPrelimit - (hne : s.Nonempty) (hbdd : BddAbove s) (hlim : ¬ IsSuccPrelimit (sSup s)) : - sSup s ∈ s := by + (hne : s.Nonempty) (hbdd : BddAbove s) (hlim : ¬ IsSuccPrelimit (sSup s)) : sSup s ∈ s := by obtain ⟨y, hy⟩ := not_forall_not.mp hlim obtain ⟨i, his, hi⟩ := exists_lt_of_lt_csSup hne hy.lt exact eq_of_le_of_not_lt (le_csSup hbdd his) (hy.2 hi) ▸ his @@ -29,8 +28,7 @@ lemma csSup_mem_of_not_isSuccPrelimit alias csSup_mem_of_not_isSuccLimit := csSup_mem_of_not_isSuccPrelimit lemma csInf_mem_of_not_isPredPrelimit - (hne : s.Nonempty) (hbdd : BddBelow s) (hlim : ¬ IsPredPrelimit (sInf s)) : - sInf s ∈ s := by + (hne : s.Nonempty) (hbdd : BddBelow s) (hlim : ¬ IsPredPrelimit (sInf s)) : sInf s ∈ s := by obtain ⟨y, hy⟩ := not_forall_not.mp hlim obtain ⟨i, his, hi⟩ := exists_lt_of_csInf_lt hne hy.lt exact eq_of_le_of_not_lt (csInf_le hbdd his) (hy.2 · hi) ▸ his @@ -39,17 +37,15 @@ lemma csInf_mem_of_not_isPredPrelimit alias csInf_mem_of_not_isPredLimit := csInf_mem_of_not_isPredPrelimit lemma exists_eq_ciSup_of_not_isSuccPrelimit - (hf : BddAbove (Set.range f)) (hf' : ¬ IsSuccPrelimit (⨆ i, f i)) : - ∃ i, f i = ⨆ i, f i := - csSup_mem_of_not_isSuccPrelimit (Set.range_nonempty f) hf hf' + (hf : BddAbove (range f)) (hf' : ¬ IsSuccPrelimit (⨆ i, f i)) : ∃ i, f i = ⨆ i, f i := + csSup_mem_of_not_isSuccPrelimit (range_nonempty f) hf hf' @[deprecated exists_eq_ciSup_of_not_isSuccPrelimit (since := "2024-09-05")] alias exists_eq_ciSup_of_not_isSuccLimit := exists_eq_ciSup_of_not_isSuccPrelimit lemma exists_eq_ciInf_of_not_isPredPrelimit - (hf : BddBelow (Set.range f)) (hf' : ¬ IsPredPrelimit (⨅ i, f i)) : - ∃ i, f i = ⨅ i, f i := - csInf_mem_of_not_isPredPrelimit (Set.range_nonempty f) hf hf' + (hf : BddBelow (range f)) (hf' : ¬ IsPredPrelimit (⨅ i, f i)) : ∃ i, f i = ⨅ i, f i := + csInf_mem_of_not_isPredPrelimit (range_nonempty f) hf hf' @[deprecated exists_eq_ciInf_of_not_isPredPrelimit (since := "2024-09-05")] alias exists_eq_ciInf_of_not_isPredLimit := exists_eq_ciInf_of_not_isPredPrelimit @@ -69,15 +65,15 @@ lemma IsGLB.mem_of_nonempty_of_not_isPredPrelimit alias IsGLB.mem_of_nonempty_of_not_isPredLimit := IsGLB.mem_of_nonempty_of_not_isPredPrelimit lemma IsLUB.exists_of_nonempty_of_not_isSuccPrelimit - (hf : IsLUB (Set.range f) x) (hx : ¬ IsSuccPrelimit x) : - ∃ i, f i = x := hf.mem_of_nonempty_of_not_isSuccPrelimit (Set.range_nonempty f) hx + (hf : IsLUB (range f) x) (hx : ¬ IsSuccPrelimit x) : ∃ i, f i = x := + hf.mem_of_nonempty_of_not_isSuccPrelimit (range_nonempty f) hx @[deprecated IsLUB.exists_of_nonempty_of_not_isSuccPrelimit (since := "2024-09-05")] alias IsLUB.exists_of_nonempty_of_not_isSuccLimit := IsLUB.exists_of_nonempty_of_not_isSuccPrelimit lemma IsGLB.exists_of_nonempty_of_not_isPredPrelimit - (hf : IsGLB (Set.range f) x) (hx : ¬ IsPredPrelimit x) : - ∃ i, f i = x := hf.mem_of_nonempty_of_not_isPredPrelimit (Set.range_nonempty f) hx + (hf : IsGLB (range f) x) (hx : ¬ IsPredPrelimit x) : ∃ i, f i = x := + hf.mem_of_nonempty_of_not_isPredPrelimit (range_nonempty f) hx @[deprecated IsGLB.exists_of_nonempty_of_not_isPredPrelimit (since := "2024-09-05")] alias IsGLB.exists_of_nonempty_of_not_isPredLimit := IsGLB.exists_of_nonempty_of_not_isPredPrelimit @@ -110,9 +106,8 @@ variable [ConditionallyCompleteLinearOrderBot α] {f : ι → α} {s : Set α} { /-- See `csSup_mem_of_not_isSuccPrelimit` for the `ConditionallyCompleteLinearOrder` version. -/ lemma csSup_mem_of_not_isSuccPrelimit' - (hbdd : BddAbove s) (hlim : ¬ IsSuccPrelimit (sSup s)) : - sSup s ∈ s := by - obtain (rfl|hs) := s.eq_empty_or_nonempty + (hbdd : BddAbove s) (hlim : ¬ IsSuccPrelimit (sSup s)) : sSup s ∈ s := by + obtain rfl | hs := s.eq_empty_or_nonempty · simp [isSuccPrelimit_bot] at hlim · exact csSup_mem_of_not_isSuccPrelimit hs hbdd hlim @@ -122,24 +117,23 @@ alias csSup_mem_of_not_isSuccLimit' := csSup_mem_of_not_isSuccPrelimit' /-- See `exists_eq_ciSup_of_not_isSuccPrelimit` for the `ConditionallyCompleteLinearOrder` version. -/ lemma exists_eq_ciSup_of_not_isSuccPrelimit' - (hf : BddAbove (Set.range f)) (hf' : ¬ IsSuccPrelimit (⨆ i, f i)) : - ∃ i, f i = ⨆ i, f i := + (hf : BddAbove (range f)) (hf' : ¬ IsSuccPrelimit (⨆ i, f i)) : ∃ i, f i = ⨆ i, f i := csSup_mem_of_not_isSuccPrelimit' hf hf' @[deprecated exists_eq_ciSup_of_not_isSuccPrelimit' (since := "2024-09-05")] alias exists_eq_ciSup_of_not_isSuccLimit' := exists_eq_ciSup_of_not_isSuccPrelimit' -lemma IsLUB.mem_of_not_isSuccPrelimit (hs : IsLUB s x) (hx : ¬ IsSuccPrelimit x) : - x ∈ s := by - obtain (rfl|hs') := s.eq_empty_or_nonempty +lemma IsLUB.mem_of_not_isSuccPrelimit (hs : IsLUB s x) (hx : ¬ IsSuccPrelimit x) : x ∈ s := by + obtain rfl | hs' := s.eq_empty_or_nonempty · simp [show x = ⊥ by simpa using hs, isSuccPrelimit_bot] at hx · exact hs.mem_of_nonempty_of_not_isSuccPrelimit hs' hx @[deprecated IsLUB.mem_of_not_isSuccPrelimit (since := "2024-09-05")] alias IsLUB.mem_of_not_isSuccLimit := IsLUB.mem_of_not_isSuccPrelimit -lemma IsLUB.exists_of_not_isSuccPrelimit (hf : IsLUB (Set.range f) x) (hx : ¬ IsSuccPrelimit x) : - ∃ i, f i = x := hf.mem_of_not_isSuccPrelimit hx +lemma IsLUB.exists_of_not_isSuccPrelimit (hf : IsLUB (range f) x) (hx : ¬ IsSuccPrelimit x) : + ∃ i, f i = x := + hf.mem_of_not_isSuccPrelimit hx @[deprecated IsLUB.exists_of_not_isSuccPrelimit (since := "2024-09-05")] alias IsLUB.exists_of_not_isSuccLimit := IsLUB.exists_of_not_isSuccPrelimit @@ -175,8 +169,7 @@ end ConditionallyCompleteLinearOrderBot section CompleteLinearOrder variable [CompleteLinearOrder α] {s : Set α} {f : ι → α} {x : α} -lemma sSup_mem_of_not_isSuccPrelimit (hlim : ¬ IsSuccPrelimit (sSup s)) : - sSup s ∈ s := by +lemma sSup_mem_of_not_isSuccPrelimit (hlim : ¬ IsSuccPrelimit (sSup s)) : sSup s ∈ s := by obtain ⟨y, hy⟩ := not_forall_not.mp hlim obtain ⟨i, his, hi⟩ := lt_sSup_iff.mp hy.lt exact eq_of_le_of_not_lt (le_sSup his) (hy.2 hi) ▸ his @@ -184,8 +177,7 @@ lemma sSup_mem_of_not_isSuccPrelimit (hlim : ¬ IsSuccPrelimit (sSup s)) : @[deprecated sSup_mem_of_not_isSuccPrelimit (since := "2024-09-05")] alias sSup_mem_of_not_isSuccLimit := sSup_mem_of_not_isSuccPrelimit -lemma sInf_mem_of_not_isPredPrelimit (hlim : ¬ IsPredPrelimit (sInf s)) : - sInf s ∈ s := by +lemma sInf_mem_of_not_isPredPrelimit (hlim : ¬ IsPredPrelimit (sInf s)) : sInf s ∈ s := by obtain ⟨y, hy⟩ := not_forall_not.mp hlim obtain ⟨i, his, hi⟩ := sInf_lt_iff.mp hy.lt exact eq_of_le_of_not_lt (sInf_le his) (hy.2 · hi) ▸ his @@ -207,15 +199,15 @@ lemma exists_eq_iInf_of_not_isPredPrelimit (hf : ¬ IsPredPrelimit (⨅ i, f i)) @[deprecated exists_eq_iInf_of_not_isPredPrelimit (since := "2024-09-05")] alias exists_eq_iInf_of_not_isPredLimit := exists_eq_iInf_of_not_isPredPrelimit -lemma IsGLB.mem_of_not_isPredPrelimit (hs : IsGLB s x) (hx : ¬ IsPredPrelimit x) : - x ∈ s := +lemma IsGLB.mem_of_not_isPredPrelimit (hs : IsGLB s x) (hx : ¬ IsPredPrelimit x) : x ∈ s := hs.sInf_eq ▸ sInf_mem_of_not_isPredPrelimit (hs.sInf_eq ▸ hx) @[deprecated IsGLB.mem_of_not_isPredPrelimit (since := "2024-09-05")] alias IsGLB.mem_of_not_isPredLimit := IsGLB.mem_of_not_isPredPrelimit -lemma IsGLB.exists_of_not_isPredPrelimit (hf : IsGLB (Set.range f) x) (hx : ¬ IsPredPrelimit x) : - ∃ i, f i = x := hf.mem_of_not_isPredPrelimit hx +lemma IsGLB.exists_of_not_isPredPrelimit (hf : IsGLB (range f) x) (hx : ¬ IsPredPrelimit x) : + ∃ i, f i = x := + hf.mem_of_not_isPredPrelimit hx @[deprecated IsGLB.exists_of_not_isPredPrelimit (since := "2024-09-05")] alias IsGLB.exists_of_not_isPredLimit := IsGLB.exists_of_not_isPredPrelimit From 878286551365eb1af02bcc97996023251716f548 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Iv=C3=A1n=20Renison?= <85908989+IvanRenison@users.noreply.github.com> Date: Tue, 10 Dec 2024 02:18:08 +0000 Subject: [PATCH 613/829] feat(Data/Fin): Add lemmas `Fin.odd_add_one_iff_even` and `Fin.even_add_one_iff_odd` (#19815) --- Mathlib/Data/Fin/Parity.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/Mathlib/Data/Fin/Parity.lean b/Mathlib/Data/Fin/Parity.lean index bd4b7b0b71e3b..2a2a76ca7d8b1 100644 --- a/Mathlib/Data/Fin/Parity.lean +++ b/Mathlib/Data/Fin/Parity.lean @@ -94,4 +94,10 @@ lemma not_even_iff_odd_of_even {n : ℕ} [NeZero n] (hn : Even n) {k : Fin n} : rw [even_iff_of_even hn, odd_iff_of_even hn] exact Nat.not_even_iff_odd +lemma odd_add_one_iff_even {n : ℕ} [NeZero n] {k : Fin n} : Odd (k + 1) ↔ Even k := + ⟨fun ⟨k, hk⟩ ↦ add_right_cancel hk ▸ even_two_mul k, Even.add_one⟩ + +lemma even_add_one_iff_odd {n : ℕ} [NeZero n] {k : Fin n} : Even (k + 1) ↔ Odd k := + ⟨fun ⟨k, hk⟩ ↦ eq_sub_iff_add_eq.mpr hk ▸ (even_add_self k).sub_odd odd_one, Odd.add_one⟩ + end Fin From 742bb01d7681a8d72da55716488a84baa750ae1f Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Tue, 10 Dec 2024 02:18:09 +0000 Subject: [PATCH 614/829] refactor(RingTheory/IntegralClosure/IntegralRestrict): Generalize from `FiniteDimensional` to `Algebra.IsAlgebraic` (#19818) This PR generalizes a few results from `FiniteDimensional` to `Algebra.IsAlgebraic`. --- .../RingTheory/IntegralClosure/IntegralRestrict.lean | 10 ++++++---- 1 file changed, 6 insertions(+), 4 deletions(-) diff --git a/Mathlib/RingTheory/IntegralClosure/IntegralRestrict.lean b/Mathlib/RingTheory/IntegralClosure/IntegralRestrict.lean index 1c84755bf4e23..b820b75229efd 100644 --- a/Mathlib/RingTheory/IntegralClosure/IntegralRestrict.lean +++ b/Mathlib/RingTheory/IntegralClosure/IntegralRestrict.lean @@ -28,10 +28,12 @@ open nonZeroDivisors variable (A K L B : Type*) [CommRing A] [CommRing B] [Algebra A B] [Field K] [Field L] [Algebra A K] [IsFractionRing A K] [Algebra B L] [Algebra K L] [Algebra A L] [IsScalarTower A B L] [IsScalarTower A K L] - [IsIntegralClosure B A L] [FiniteDimensional K L] + [IsIntegralClosure B A L] section galois +variable [Algebra.IsAlgebraic K L] + /-- The lift `End(B/A) → End(L/K)` in an ALKB setup. This is inverse to the restriction. See `galRestrictHom`. -/ noncomputable @@ -114,7 +116,9 @@ lemma algebraMap_galRestrict_apply (σ : L ≃ₐ[K] L) (x : B) : algebraMap B L (galRestrict A K L B σ x) = σ (algebraMap B L x) := algebraMap_galRestrictHom_apply A K L B σ.toAlgHom x -variable (K L B) +end galois + +variable [FiniteDimensional K L] lemma prod_galRestrict_eq_norm [IsGalois K L] [IsIntegrallyClosed A] (x : B) : (∏ σ : L ≃ₐ[K] L, galRestrict A K L B σ x) = @@ -125,8 +129,6 @@ lemma prod_galRestrict_eq_norm [IsGalois K L] [IsIntegrallyClosed A] (x : B) : simp only [map_prod, algebraMap_galRestrict_apply, IsIntegralClosure.algebraMap_mk', Algebra.norm_eq_prod_automorphisms, AlgHom.coe_coe, RingHom.coe_comp, Function.comp_apply] -end galois - attribute [local instance] FractionRing.liftAlgebra FractionRing.isScalarTower_liftAlgebra noncomputable From b2dc853962c8e833069e068f19de5854d55a9be8 Mon Sep 17 00:00:00 2001 From: Andrew Yang <36414270+erdOne@users.noreply.github.com> Date: Tue, 10 Dec 2024 02:18:11 +0000 Subject: [PATCH 615/829] chore(RingTheory/LocalRing): remove diamond (#19819) --- Mathlib/RingTheory/LocalRing/ResidueField/Basic.lean | 4 ---- 1 file changed, 4 deletions(-) diff --git a/Mathlib/RingTheory/LocalRing/ResidueField/Basic.lean b/Mathlib/RingTheory/LocalRing/ResidueField/Basic.lean index ad680a8c1d695..cf35586f78472 100644 --- a/Mathlib/RingTheory/LocalRing/ResidueField/Basic.lean +++ b/Mathlib/RingTheory/LocalRing/ResidueField/Basic.lean @@ -164,9 +164,6 @@ variable [Algebra R S] [IsLocalHom (algebraMap R S)] noncomputable instance : Algebra (ResidueField R) (ResidueField S) := (ResidueField.map (algebraMap R S)).toAlgebra -noncomputable instance : Algebra R (ResidueField S) := - ((ResidueField.map <| algebraMap R S).comp <| residue R).toAlgebra - instance : IsScalarTower R (ResidueField R) (ResidueField S) := IsScalarTower.of_algebraMap_eq (congrFun rfl) @@ -176,7 +173,6 @@ instance finiteDimensional_of_noetherian [IsNoetherian R S] : isNoetherian_of_tower R (S := ResidueField R) (M := ResidueField S) _ convert isNoetherian_of_surjective S (Ideal.Quotient.mkₐ R (maximalIdeal S)).toLinearMap (LinearMap.range_eq_top.mpr Ideal.Quotient.mk_surjective) - exact Algebra.algebra_ext _ _ (fun r => rfl) -- We want to be able to refer to `hfin` set_option linter.unusedVariables false in From 4dfbb55f0c2ba2ac088e8243b259908358110e83 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Pol=27tta=20/=20Miyahara=20K=C5=8D?= Date: Tue, 10 Dec 2024 02:18:12 +0000 Subject: [PATCH 616/829] chore(Mathlib/SetTheory/ZFC/Basic): `pp_with_univ` for `ZFSet` & related (#19824) `pp_with_univ` to `ZFSet`, `PSet` and `Class` --- Mathlib/SetTheory/ZFC/Basic.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/SetTheory/ZFC/Basic.lean b/Mathlib/SetTheory/ZFC/Basic.lean index 00484448b82b2..535579e53318c 100644 --- a/Mathlib/SetTheory/ZFC/Basic.lean +++ b/Mathlib/SetTheory/ZFC/Basic.lean @@ -57,6 +57,7 @@ open Function (OfArity) is a family of pre-sets indexed by a type in `Type u`. The ZFC universe is defined as a quotient of this to ensure extensionality. -/ +@[pp_with_univ] inductive PSet : Type (u + 1) | mk (α : Type u) (A : α → PSet) : PSet @@ -515,6 +516,7 @@ end PSet /-- The ZFC universe of sets consists of the type of pre-sets, quotiented by extensional equivalence. -/ +@[pp_with_univ] def ZFSet : Type (u + 1) := Quotient PSet.setoid.{u} @@ -1346,6 +1348,7 @@ end ZFSet We define `Class` as `Set ZFSet`, as this allows us to get many instances automatically. However, in practice, we treat it as (the definitionally equal) `ZFSet → Prop`. This means, the preferred way to state that `x : ZFSet` belongs to `A : Class` is to write `A x`. -/ +@[pp_with_univ] def Class := Set ZFSet deriving HasSubset, EmptyCollection, Nonempty, Union, Inter, HasCompl, SDiff From 0689cbcda8cb172a0aa55a6b31fa5b65eaa45e2b Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Tue, 10 Dec 2024 02:18:13 +0000 Subject: [PATCH 617/829] chore: fix some file references (#19830) --- Mathlib/Algebra/Module/Submodule/Invariant.lean | 2 +- Mathlib/Analysis/Calculus/FDeriv/Basic.lean | 2 +- Mathlib/Analysis/Calculus/TangentCone.lean | 12 +++++++----- Mathlib/CategoryTheory/Abelian/Refinements.lean | 6 +++--- Mathlib/CategoryTheory/Closed/Enrichment.lean | 4 ++-- Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean | 2 +- Mathlib/CategoryTheory/Monoidal/Functor.lean | 4 ++-- Mathlib/Data/Finset/Lattice/Fold.lean | 4 ++-- Mathlib/Data/Nat/Prime/Int.lean | 2 +- Mathlib/Data/Rat/Denumerable.lean | 2 +- Mathlib/Data/Setoid/Basic.lean | 2 +- Mathlib/Data/Subtype.lean | 2 +- Mathlib/Logic/Equiv/Basic.lean | 4 ++-- Mathlib/Logic/Equiv/Defs.lean | 9 +++++---- Mathlib/MeasureTheory/Group/Prod.lean | 2 +- Mathlib/MeasureTheory/Measure/Haar/Basic.lean | 4 ++-- Mathlib/MeasureTheory/Measure/Haar/Unique.lean | 2 +- Mathlib/Order/CompleteLattice/Finset.lean | 3 ++- Mathlib/Topology/Defs/Basic.lean | 2 +- 19 files changed, 37 insertions(+), 33 deletions(-) diff --git a/Mathlib/Algebra/Module/Submodule/Invariant.lean b/Mathlib/Algebra/Module/Submodule/Invariant.lean index 0e8d20eb33819..cfb3ce8b980af 100644 --- a/Mathlib/Algebra/Module/Submodule/Invariant.lean +++ b/Mathlib/Algebra/Module/Submodule/Invariant.lean @@ -13,7 +13,7 @@ In this file we defined the type `Module.End.invtSubmodule`, associated to a lin a module. Its utilty stems primarily from those occasions on which we wish to take advantage of the lattice structure of invariant submodules. -See also `Module.AEval`. +See also `Mathlib.Algebra.Polynomial.Module.AEval`. -/ diff --git a/Mathlib/Analysis/Calculus/FDeriv/Basic.lean b/Mathlib/Analysis/Calculus/FDeriv/Basic.lean index b396b208a0061..c381a83392646 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Basic.lean @@ -70,7 +70,7 @@ example (x : ℝ) (h : 1 + sin x ≠ 0) : DifferentiableAt ℝ (fun x ↦ exp x simp [h] ``` Of course, these examples only work once `exp`, `cos` and `sin` have been shown to be -differentiable, in `Analysis.SpecialFunctions.Trigonometric`. +differentiable, in `Mathlib.Analysis.SpecialFunctions.Trigonometric.Deriv`. The simplifier is not set up to compute the Fréchet derivative of maps (as these are in general complicated multidimensional linear maps), but it will compute one-dimensional derivatives, diff --git a/Mathlib/Analysis/Calculus/TangentCone.lean b/Mathlib/Analysis/Calculus/TangentCone.lean index f3ac765b70a30..c818f6cfc6e58 100644 --- a/Mathlib/Analysis/Calculus/TangentCone.lean +++ b/Mathlib/Analysis/Calculus/TangentCone.lean @@ -25,9 +25,10 @@ of the derivative. This is why their names reflect their uses, and not how they ## Implementation details -Note that this file is imported by `Fderiv.Basic`. Hence, derivatives are not defined yet. The -property of uniqueness of the derivative is therefore proved in `Fderiv.Basic`, but based on the -properties of the tangent cone we prove here. +Note that this file is imported by `Mathlib.Analysis.Calculus.FDeriv.Basic`. Hence, derivatives are +not defined yet. The property of uniqueness of the derivative is therefore proved in +`Mathlib.Analysis.Calculus.FDeriv.Basic`, but based on the properties of the tangent cone we prove +here. -/ @@ -50,7 +51,8 @@ def tangentConeAt (s : Set E) (x : E) : Set E := /-- A property ensuring that the tangent cone to `s` at `x` spans a dense subset of the whole space. The main role of this property is to ensure that the differential within `s` at `x` is unique, -hence this name. The uniqueness it asserts is proved in `UniqueDiffWithinAt.eq` in `Fderiv.Basic`. +hence this name. The uniqueness it asserts is proved in `UniqueDiffWithinAt.eq` in +`Mathlib.Analysis.Calculus.FDeriv.Basic`. To avoid pathologies in dimension 0, we also require that `x` belongs to the closure of `s` (which is automatic when `E` is not `0`-dimensional). -/ @[mk_iff] @@ -61,7 +63,7 @@ structure UniqueDiffWithinAt (s : Set E) (x : E) : Prop where /-- A property ensuring that the tangent cone to `s` at any of its points spans a dense subset of the whole space. The main role of this property is to ensure that the differential along `s` is unique, hence this name. The uniqueness it asserts is proved in `UniqueDiffOn.eq` in -`Fderiv.Basic`. -/ +`Mathlib.Analysis.Calculus.FDeriv.Basic`. -/ def UniqueDiffOn (s : Set E) : Prop := ∀ x ∈ s, UniqueDiffWithinAt 𝕜 s x diff --git a/Mathlib/CategoryTheory/Abelian/Refinements.lean b/Mathlib/CategoryTheory/Abelian/Refinements.lean index 8dffabb705f3f..75582a7bcc024 100644 --- a/Mathlib/CategoryTheory/Abelian/Refinements.lean +++ b/Mathlib/CategoryTheory/Abelian/Refinements.lean @@ -44,17 +44,17 @@ category is exact if and only if it is exact up to refinements (see `ShortComplex.exact_iff_exact_up_to_refinements`). As it is outlined in the documentation of the file -`CategoryTheory.Abelian.Pseudoelements`, the Freyd-Mitchell +`Mathlib.CategoryTheory.Abelian.Pseudoelements`, the Freyd-Mitchell embedding theorem implies the existence of a faithful and exact functor `ι` from an abelian category `C` to the category of abelian groups. If we define a pseudo-element of `X : C` to be an element in `ι.obj X`, one may do diagram chases in any abelian category using these pseudo-elements. However, using this approach would require proving this embedding theorem! Currently, mathlib contains a weaker notion of pseudo-elements -`CategoryTheory.Abelian.Pseudoelements`. Some theorems can be obtained +`Mathlib.CategoryTheory.Abelian.Pseudoelements`. Some theorems can be obtained using this notion, but there is the issue that for this notion of pseudo-elements a morphism `X ⟶ Y` in `C` is not determined by -its action on pseudo-elements (see also `Counterexamples/Pseudoelement`). +its action on pseudo-elements (see also `Counterexamples/Pseudoelement.lean`). On the contrary, the approach consisting of working up to refinements does not require the introduction of other types: we only need to work with morphisms `A ⟶ X` in `C` which we may consider as being diff --git a/Mathlib/CategoryTheory/Closed/Enrichment.lean b/Mathlib/CategoryTheory/Closed/Enrichment.lean index 1d263f92528ed..d0cedeb852160 100644 --- a/Mathlib/CategoryTheory/Closed/Enrichment.lean +++ b/Mathlib/CategoryTheory/Closed/Enrichment.lean @@ -15,9 +15,9 @@ where the hom-object is given by the internal hom (coming from the closed struct We use `scoped instance` to avoid potential issues where `C` may also have a `C`-category structure coming from another source (e.g. the type of simplicial sets `SSet.{v}` has an instance of `EnrichedCategory SSet.{v}` as a category of simplicial objects; -see `AlgebraicTopology/SimplicialCategory/SimplicialObject`). +see `Mathlib/AlgebraicTopology/SimplicialCategory/SimplicialObject.lean`). -All structure field values are defined in `Closed/Monoidal`. +All structure field values are defined in `Mathlib/CategoryTheory/Closed/Monoidal.lean`. -/ diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean b/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean index 7861c1e682600..6c02b797dda67 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean @@ -23,7 +23,7 @@ pointwise. ## See also There is also a relative version of this statement where `F : J ⥤ Over A` for some presheaf -`A`, see `CategoryTheory.Comma.Presheaf.Colimit`. +`A`, see `Mathlib.CategoryTheory.Comma.Presheaf.Colimit`. -/ diff --git a/Mathlib/CategoryTheory/Monoidal/Functor.lean b/Mathlib/CategoryTheory/Monoidal/Functor.lean index 72a7a69ef0039..a2388f08a6793 100644 --- a/Mathlib/CategoryTheory/Monoidal/Functor.lean +++ b/Mathlib/CategoryTheory/Monoidal/Functor.lean @@ -25,9 +25,9 @@ inverse isomorphisms. We show that the composition of (lax) monoidal functors gives a (lax) monoidal functor. -See `CategoryTheory.Monoidal.NaturalTransformation` for monoidal natural transformations. +See `Mathlib.CategoryTheory.Monoidal.NaturalTransformation` for monoidal natural transformations. -We show in `CategoryTheory.Monoidal.Mon_` that lax monoidal functors take monoid objects +We show in `Mathlib.CategoryTheory.Monoidal.Mon_` that lax monoidal functors take monoid objects to monoid objects. ## References diff --git a/Mathlib/Data/Finset/Lattice/Fold.lean b/Mathlib/Data/Finset/Lattice/Fold.lean index 97a141e5d4d22..506bb676f71b9 100644 --- a/Mathlib/Data/Finset/Lattice/Fold.lean +++ b/Mathlib/Data/Finset/Lattice/Fold.lean @@ -18,8 +18,8 @@ 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. +See also `Mathlib/Order/CompleteLattice/Finset.lean`, which is instead concerned with how big +lattice or set operations behave when indexed by a finset. -/ assert_not_exists OrderedCommMonoid diff --git a/Mathlib/Data/Nat/Prime/Int.lean b/Mathlib/Data/Nat/Prime/Int.lean index f4fc6f0359bbe..d357a3d303050 100644 --- a/Mathlib/Data/Nat/Prime/Int.lean +++ b/Mathlib/Data/Nat/Prime/Int.lean @@ -9,7 +9,7 @@ import Mathlib.Data.Nat.Prime.Basic /-! # Prime numbers in the naturals and the integers -TODO: This file can probably be merged with `Data/Nat/Int/NatPrime.lean`. +TODO: This file can probably be merged with `Mathlib/Data/Int/NatPrime.lean`. -/ diff --git a/Mathlib/Data/Rat/Denumerable.lean b/Mathlib/Data/Rat/Denumerable.lean index 63a9a4318fa46..116e03e72a7dd 100644 --- a/Mathlib/Data/Rat/Denumerable.lean +++ b/Mathlib/Data/Rat/Denumerable.lean @@ -13,7 +13,7 @@ import Mathlib.Logic.Denumerable This file proves that ℚ is denumerable. -The fact that ℚ has cardinality ℵ₀ is proved in `Data.Rat.Cardinal` +The fact that ℚ has cardinality ℵ₀ is proved in `Mathlib.Data.Rat.Cardinal` -/ assert_not_exists Module diff --git a/Mathlib/Data/Setoid/Basic.lean b/Mathlib/Data/Setoid/Basic.lean index d860f84f98bd6..3599ac45fefc2 100644 --- a/Mathlib/Data/Setoid/Basic.lean +++ b/Mathlib/Data/Setoid/Basic.lean @@ -19,7 +19,7 @@ The complete lattice instance for equivalence relations could have been defined the Galois insertion of equivalence relations on α into binary relations on α, and then using `CompleteLattice.copy` to define a complete lattice instance with more appropriate definitional equalities (a similar example is `Filter.CompleteLattice` in -`Order/Filter/Basic.lean`). This does not save space, however, and is less clear. +`Mathlib/Order/Filter/Basic.lean`). This does not save space, however, and is less clear. Partitions are not defined as a separate structure here; users are encouraged to reason about them using the existing `Setoid` and its infrastructure. diff --git a/Mathlib/Data/Subtype.lean b/Mathlib/Data/Subtype.lean index e343f27f48c4b..79c2b5a26c935 100644 --- a/Mathlib/Data/Subtype.lean +++ b/Mathlib/Data/Subtype.lean @@ -33,7 +33,7 @@ attribute [coe] Subtype.val initialize_simps_projections Subtype (val → coe) /-- A version of `x.property` or `x.2` where `p` is syntactically applied to the coercion of `x` - instead of `x.1`. A similar result is `Subtype.mem` in `Data.Set.Basic`. -/ + instead of `x.1`. A similar result is `Subtype.mem` in `Mathlib.Data.Set.Basic`. -/ theorem prop (x : Subtype p) : p x := x.2 diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 070edf6a55b4e..0a2b42804e59c 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -22,7 +22,7 @@ import Mathlib.Tactic.CC /-! # Equivalence between types -In this file we continue the work on equivalences begun in `Logic/Equiv/Defs.lean`, defining +In this file we continue the work on equivalences begun in `Mathlib/Logic/Equiv/Defs.lean`, defining * canonical isomorphisms between various types: e.g., @@ -38,7 +38,7 @@ In this file we continue the work on equivalences begun in `Logic/Equiv/Defs.lea `eb : β₁ ≃ β₂` using `Prod.map`. More definitions of this kind can be found in other files. - E.g., `Logic/Equiv/TransferInstance.lean` does it for many algebraic type classes like + E.g., `Mathlib/Logic/Equiv/TransferInstance.lean` does it for many algebraic type classes like `Group`, `Module`, etc. ## Tags diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index 9c482264afc23..ceaaf86181c2d 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -19,7 +19,7 @@ In this file we define two types: not equality!) to express that various `Type`s or `Sort`s are equivalent. * `Equiv.Perm α`: the group of permutations `α ≃ α`. More lemmas about `Equiv.Perm` can be found in - `GroupTheory.Perm`. + `Mathlib.GroupTheory.Perm`. Then we define @@ -41,10 +41,11 @@ Then we define - `Equiv.unique` takes `e : α ≃ β` and `[Unique β]` and returns `Unique α`; - `Equiv.decidableEq` takes `e : α ≃ β` and `[DecidableEq β]` and returns `DecidableEq α`. - More definitions of this kind can be found in other files. E.g., `Data.Equiv.TransferInstance` - does it for many algebraic type classes like `Group`, `Module`, etc. + More definitions of this kind can be found in other files. + E.g., `Mathlib.Logic.Equiv.TransferInstance` does it for many algebraic type classes like + `Group`, `Module`, etc. -Many more such isomorphisms and operations are defined in `Logic.Equiv.Basic`. +Many more such isomorphisms and operations are defined in `Mathlib.Logic.Equiv.Basic`. ## Tags diff --git a/Mathlib/MeasureTheory/Group/Prod.lean b/Mathlib/MeasureTheory/Group/Prod.lean index a33b49bf00b01..db452a55297de 100644 --- a/Mathlib/MeasureTheory/Group/Prod.lean +++ b/Mathlib/MeasureTheory/Group/Prod.lean @@ -34,7 +34,7 @@ Note that this theory only applies in measurable groups, i.e., when multiplicati are measurable. This is not the case in general in locally compact groups, or even in compact groups, when the topology is not second-countable. For arguments along the same line, but using continuous functions instead of measurable sets and working in the general locally compact -setting, see the file `MeasureTheory.Measure.Haar.Unique.lean`. +setting, see the file `Mathlib/MeasureTheory/Measure/Haar/Unique.lean`. -/ diff --git a/Mathlib/MeasureTheory/Measure/Haar/Basic.lean b/Mathlib/MeasureTheory/Measure/Haar/Basic.lean index 546c35be2dda7..caed697da9d50 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Basic.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Basic.lean @@ -34,7 +34,7 @@ where `ᵒ` denotes the interior. We also give a form of uniqueness of Haar measure, for σ-finite measures on second-countable locally compact groups. For more involved statements not assuming second-countability, see -the file `MeasureTheory.Measure.Haar.Unique`. +the file `Mathlib/MeasureTheory/Measure/Haar/Unique.lean`. ## Main Declarations @@ -626,7 +626,7 @@ bypass all topological arguments, and conclude using uniqueness of σ-finite lef in measurable groups. For more general uniqueness statements without second-countability assumptions, -see the file `MeasureTheory.Measure.Haar.Unique`. +see the file `Mathlib/MeasureTheory/Measure/Haar/Unique.lean`. -/ variable [SecondCountableTopology G] diff --git a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean index a8474fe9a8a32..5d1a0d0477fb8 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean @@ -58,7 +58,7 @@ measure. We follow McQuillan's answer at https://mathoverflow.net/questions/4566 On second-countable groups, one can arrive to slightly different uniqueness results by using that the operations are measurable. In particular, one can get uniqueness assuming σ-finiteness of the measures but discarding the assumption that they are finite on compact sets. See -`haarMeasure_unique` in the file `MeasureTheory.Measure.Haar.Basic`. +`haarMeasure_unique` in the file `Mathlib/MeasureTheory/Measure/Haar/Basic.lean`. ## References diff --git a/Mathlib/Order/CompleteLattice/Finset.lean b/Mathlib/Order/CompleteLattice/Finset.lean index b02e5eb918efa..b3607813b1f92 100644 --- a/Mathlib/Order/CompleteLattice/Finset.lean +++ b/Mathlib/Order/CompleteLattice/Finset.lean @@ -11,7 +11,8 @@ import Mathlib.Data.Set.Lattice 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. +See also `Mathlib/Data/Finset/Lattice.lean`, which is concerned with folding binary lattice +operations over a finset. -/ assert_not_exists OrderedCommMonoid diff --git a/Mathlib/Topology/Defs/Basic.lean b/Mathlib/Topology/Defs/Basic.lean index e7526dd6e9024..ba372a4470a09 100644 --- a/Mathlib/Topology/Defs/Basic.lean +++ b/Mathlib/Topology/Defs/Basic.lean @@ -100,7 +100,7 @@ def IsClopen (s : Set X) : Prop := /-- A set is locally closed if it is the intersection of some open set and some closed set. -Also see `isLocallyClosed_tfae` and other lemmas in `Mathlib/Topology/LocallyClosed`. +Also see `isLocallyClosed_tfae` and other lemmas in `Mathlib/Topology/LocallyClosed.lean`. -/ def IsLocallyClosed (s : Set X) : Prop := ∃ (U Z : Set X), IsOpen U ∧ IsClosed Z ∧ s = U ∩ Z From 8544aa528cc43b29b7bd4f71e9b21b944d15fda4 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Tue, 10 Dec 2024 02:18:15 +0000 Subject: [PATCH 618/829] chore: fix some notation commands (#19844) * These notations were not used in the pretty-printer. Now they are. * Exception: The notation in `Mathlib/Algebra/Category/Grp/EpiMono.lean` still doesn't pretty-print correctly. This is still an improvement, because the previous command didn't declare the notation correctly, and it was unused in the rest of the file. * There is more to do (everything with `set_option quotPrecheck false` is suspicious). --- Mathlib/Algebra/Category/Grp/EpiMono.lean | 7 +++---- Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean | 10 ++++------ Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean | 3 +-- Mathlib/Geometry/Manifold/Diffeomorph.lean | 2 +- Mathlib/Geometry/Manifold/Instances/Real.lean | 4 ++-- Mathlib/MeasureTheory/VectorMeasure/Basic.lean | 3 +-- Mathlib/RingTheory/WittVector/Isocrystal.lean | 4 ++-- 7 files changed, 14 insertions(+), 19 deletions(-) diff --git a/Mathlib/Algebra/Category/Grp/EpiMono.lean b/Mathlib/Algebra/Category/Grp/EpiMono.lean index 308229317d4be..b1252eb15d7cd 100644 --- a/Mathlib/Algebra/Category/Grp/EpiMono.lean +++ b/Mathlib/Algebra/Category/Grp/EpiMono.lean @@ -88,13 +88,12 @@ theorem mono_iff_injective : Mono f ↔ Function.Injective f := namespace SurjectiveOfEpiAuxs -set_option quotPrecheck false in -local notation "X" => Set.range (· • (f.range : Set B) : B → Set B) +local notation3 "X" => Set.range (· • (f.range : Set B) : B → Set B) /-- Define `X'` to be the set of all left cosets with an extra point at "infinity". -/ inductive XWithInfinity - | fromCoset : Set.range (· • (f.range : Set B) : B → Set B) → XWithInfinity + | fromCoset : X → XWithInfinity | infinity : XWithInfinity open XWithInfinity Equiv.Perm @@ -222,7 +221,7 @@ The strategy is the following: assuming `epi f` -/ -theorem g_apply_fromCoset (x : B) (y : Set.range (· • (f.range : Set B) : B → Set B)) : +theorem g_apply_fromCoset (x : B) (y : X) : g x (fromCoset y) = fromCoset ⟨x • ↑y, by obtain ⟨z, hz⟩ := y.2; exact ⟨x * z, by simp [← hz, smul_smul]⟩⟩ := rfl diff --git a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean index 3a29eb4977c1c..dc2a61ca1a817 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean @@ -89,13 +89,11 @@ open Function variable {M₀ G₀ : Type*} (α : Type*) -set_option quotPrecheck false in -/-- Local notation for the nonnegative elements of a type `α`. TODO: actually make local. -/ -notation "α≥0" => { x : α // 0 ≤ x } +/-- Local notation for the nonnegative elements of a type `α`. -/ +local notation3 "α≥0" => { x : α // 0 ≤ x } -set_option quotPrecheck false in -/-- Local notation for the positive elements of a type `α`. TODO: actually make local. -/ -notation "α>0" => { x : α // 0 < x } +/-- Local notation for the positive elements of a type `α`. -/ +local notation3 "α>0" => { x : α // 0 < x } section Abbreviations diff --git a/Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean b/Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean index c7ee1ad189a72..777aac3b212a6 100644 --- a/Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean +++ b/Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean @@ -484,10 +484,9 @@ instance : Category (CosimplicialObject C) := by namespace CosimplicialObject -set_option quotPrecheck false in /-- `X _[n]` denotes the `n`th-term of the cosimplicial object X -/ scoped[Simplicial] - notation:1000 X " _[" n "]" => + notation3:1000 X " _[" n "]" => (X : CategoryTheory.CosimplicialObject _).obj (SimplexCategory.mk n) instance {J : Type v} [SmallCategory J] [HasLimitsOfShape J C] : diff --git a/Mathlib/Geometry/Manifold/Diffeomorph.lean b/Mathlib/Geometry/Manifold/Diffeomorph.lean index ced591a0a2f5f..ed39ae6e822cd 100644 --- a/Mathlib/Geometry/Manifold/Diffeomorph.lean +++ b/Mathlib/Geometry/Manifold/Diffeomorph.lean @@ -83,7 +83,7 @@ scoped[Manifold] notation E " ≃ₘ^" n:1000 "[" 𝕜 "] " E' => Diffeomorph /-- Infinitely differentiable diffeomorphism between `E` and `E'`. -/ scoped[Manifold] - notation E " ≃ₘ[" 𝕜 "] " E' => + notation3 E " ≃ₘ[" 𝕜 "] " E' => Diffeomorph (modelWithCornersSelf 𝕜 E) (modelWithCornersSelf 𝕜 E') E E' ⊤ namespace Diffeomorph diff --git a/Mathlib/Geometry/Manifold/Instances/Real.lean b/Mathlib/Geometry/Manifold/Instances/Real.lean index 299a57f94ff35..3bea4b00685c4 100644 --- a/Mathlib/Geometry/Manifold/Instances/Real.lean +++ b/Mathlib/Geometry/Manifold/Instances/Real.lean @@ -219,13 +219,13 @@ def modelWithCornersEuclideanQuadrant (n : ℕ) : /-- The model space used to define `n`-dimensional real manifolds without boundary. -/ scoped[Manifold] - notation "𝓡 " n => + notation3 "𝓡 " n => (modelWithCornersSelf ℝ (EuclideanSpace ℝ (Fin n)) : ModelWithCorners ℝ (EuclideanSpace ℝ (Fin n)) (EuclideanSpace ℝ (Fin n))) /-- The model space used to define `n`-dimensional real manifolds with boundary. -/ scoped[Manifold] - notation "𝓡∂ " n => + notation3 "𝓡∂ " n => (modelWithCornersEuclideanHalfSpace n : ModelWithCorners ℝ (EuclideanSpace ℝ (Fin n)) (EuclideanHalfSpace n)) diff --git a/Mathlib/MeasureTheory/VectorMeasure/Basic.lean b/Mathlib/MeasureTheory/VectorMeasure/Basic.lean index e5e7a7d55a5d6..c316fb186a98b 100644 --- a/Mathlib/MeasureTheory/VectorMeasure/Basic.lean +++ b/Mathlib/MeasureTheory/VectorMeasure/Basic.lean @@ -710,9 +710,8 @@ theorem le_iff' : v ≤ w ↔ ∀ i, v i ≤ w i := by end -set_option quotPrecheck false in -- Porting note: error message suggested to do this scoped[MeasureTheory] - notation:50 v " ≤[" i:50 "] " w:50 => + notation3:50 v " ≤[" i:50 "] " w:50 => MeasureTheory.VectorMeasure.restrict v i ≤ MeasureTheory.VectorMeasure.restrict w i section diff --git a/Mathlib/RingTheory/WittVector/Isocrystal.lean b/Mathlib/RingTheory/WittVector/Isocrystal.lean index 9fc53cafc60b5..38d771bf5257c 100644 --- a/Mathlib/RingTheory/WittVector/Isocrystal.lean +++ b/Mathlib/RingTheory/WittVector/Isocrystal.lean @@ -90,11 +90,11 @@ instance inv_pair₂ : RingHomInvPair ((FractionRing.frobenius p k).symm : K(p, RingHomInvPair.of_ringEquiv (FractionRing.frobenius p k).symm scoped[Isocrystal] - notation:50 M " →ᶠˡ[" p ", " k "] " M₂ => + notation3:50 M " →ᶠˡ[" p ", " k "] " M₂ => LinearMap (WittVector.FractionRing.frobeniusRingHom p k) M M₂ scoped[Isocrystal] - notation:50 M " ≃ᶠˡ[" p ", " k "] " M₂ => + notation3:50 M " ≃ᶠˡ[" p ", " k "] " M₂ => LinearEquiv (WittVector.FractionRing.frobeniusRingHom p k) M M₂ /-! ### Isocrystals -/ From ebd71de74db22aaa0abbfab6799a22bcecda3476 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Tue, 10 Dec 2024 02:18:16 +0000 Subject: [PATCH 619/829] doc(NumberTheory/Padics/MahlerBasis): add credit line (#19845) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The file `NumberTheory/Padics/MahlerBasis.lean` is largely based on the ETH Zürich bachelor thesis of Giulio Caflisch. This PR gives Giulio due credit for this -- he is already included in the "Authors" line, but this does not appear in the online reference manual. I am following here the example of [JacobiSum/Basic](https://leanprover-community.github.io/mathlib4_docs/Mathlib/NumberTheory/JacobiSum/Basic.html), which is likewise based on a bachelor thesis project. --- Mathlib/NumberTheory/Padics/MahlerBasis.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/NumberTheory/Padics/MahlerBasis.lean b/Mathlib/NumberTheory/Padics/MahlerBasis.lean index 8adaf56467425..827e9456b8035 100644 --- a/Mathlib/NumberTheory/Padics/MahlerBasis.lean +++ b/Mathlib/NumberTheory/Padics/MahlerBasis.lean @@ -25,6 +25,9 @@ converges (uniformly) to `f`, and this construction defines a Banach-space isomo For this, we follow the argument of Bojanić [bojanic74]. +The formalisation of Mahler's theorem presented here is based on code written by Giulio Caflisch +for his bachelor's thesis at ETH Zürich. + ## References * [R. Bojanić, *A simple proof of Mahler's theorem on approximation of continuous functions of a From 6c84ec343253af71c4e5827f2a94d2f555cba8c1 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 10 Dec 2024 02:59:30 +0000 Subject: [PATCH 620/829] feat: norm_num for Int.ceil (#19669) Long ago I tried writing this in https://github.com/leanprover-community/mathlib3/pull/16502 for Lean 3. --- Mathlib/Data/Rat/Floor.lean | 52 ++++++++++++++++++++++++++++++++++- MathlibTest/norm_num_ext.lean | 6 ++++ 2 files changed, 57 insertions(+), 1 deletion(-) diff --git a/Mathlib/Data/Rat/Floor.lean b/Mathlib/Data/Rat/Floor.lean index 7b67abd634887..a916bde474c7d 100644 --- a/Mathlib/Data/Rat/Floor.lean +++ b/Mathlib/Data/Rat/Floor.lean @@ -50,6 +50,11 @@ instance : FloorRing ℚ := protected theorem floor_def {q : ℚ} : ⌊q⌋ = q.num / q.den := Rat.floor_def' q +protected theorem ceil_def (q : ℚ) : ⌈q⌉ = -(-q.num / ↑q.den) := by + change -⌊-q⌋ = _ + rw [Rat.floor_def, num_neg_eq_neg_num, den_neg_eq_den] + + @[norm_cast] theorem floor_intCast_div_natCast (n : ℤ) (d : ℕ) : ⌊(↑n / ↑d : ℚ)⌋ = n / (↑d : ℤ) := by rw [Rat.floor_def] @@ -63,10 +68,19 @@ theorem floor_intCast_div_natCast (n : ℤ) (d : ℕ) : ⌊(↑n / ↑d : ℚ) refine (Int.mul_ediv_mul_of_pos _ _ <| pos_of_mul_pos_left ?_ <| Int.natCast_nonneg q.den).symm rwa [← d_eq_c_mul_denom, Int.natCast_pos] +@[norm_cast] +theorem ceil_intCast_div_natCast (n : ℤ) (d : ℕ) : ⌈(↑n / ↑d : ℚ)⌉ = -((-n) / (↑d : ℤ)) := by + conv_lhs => rw [← neg_neg ⌈_⌉, ← floor_neg] + rw [← neg_div, ← Int.cast_neg, floor_intCast_div_natCast] + @[norm_cast] theorem floor_natCast_div_natCast (n d : ℕ) : ⌊(↑n / ↑d : ℚ)⌋ = n / d := floor_intCast_div_natCast n d +@[norm_cast] +theorem ceil_natCast_div_natCast (n d : ℕ) : ⌈(↑n / ↑d : ℚ)⌉ = -((-n) / d) := + ceil_intCast_div_natCast n d + @[norm_cast] theorem natFloor_natCast_div_natCast (n d : ℕ) : ⌊(↑n / ↑d : ℚ)⌋₊ = n / d := by rw [← Int.ofNat_inj, Int.natCast_floor_eq_floor (by positivity)] @@ -128,10 +142,46 @@ def evalIntFloor : NormNumExt where eval {u αZ} e := do let _i ← synthInstanceQ q(LinearOrderedField $α) assertInstancesCommute have z : Q(ℤ) := mkRawIntLit ⌊q⌋ - letI : $z =Q ⌊$n / $d⌋ := ⟨⟩ + letI : $z =Q $n / $d := ⟨⟩ return .isInt q(inferInstance) z ⌊q⌋ q(isInt_intFloor_ofIsRat _ $n $d $h) | _, _, _ => failure +theorem isNat_intCeil {R} [LinearOrderedRing R] [FloorRing R] (r : R) (m : ℕ) : + IsNat r m → IsNat ⌈r⌉ m := by rintro ⟨⟨⟩⟩; exact ⟨by simp⟩ + +theorem isInt_intCeil {R} [LinearOrderedRing R] [FloorRing R] (r : R) (m : ℤ) : + IsInt r m → IsInt ⌈r⌉ m := by rintro ⟨⟨⟩⟩; exact ⟨by simp⟩ + +theorem isInt_intCeil_ofIsRat (r : α) (n : ℤ) (d : ℕ) : + IsRat r n d → IsInt ⌈r⌉ (-(-n / d)) := by + rintro ⟨inv, rfl⟩ + constructor + simp only [invOf_eq_inv, ← div_eq_mul_inv, Int.cast_id] + rw [← ceil_intCast_div_natCast n d, ← ceil_cast (α := α), Rat.cast_div, + cast_intCast, cast_natCast] + +/-- `norm_num` extension for `Int.ceil` -/ +@[norm_num ⌈_⌉] +def evalIntCeil : NormNumExt where eval {u αZ} e := do + match u, αZ, e with + | 0, ~q(ℤ), ~q(@Int.ceil $α $instR $instF $x) => + match ← derive x with + | .isBool .. => failure + | .isNat _ _ pb => do + assertInstancesCommute + return .isNat q(inferInstance) _ q(isNat_intCeil $x _ $pb) + | .isNegNat _ _ pb => do + assertInstancesCommute + -- ceil always keeps naturals negative, so we can shortcut `.isInt` + return .isNegNat q(inferInstance) _ q(isInt_intCeil _ _ $pb) + | .isRat _ q n d h => do + let _i ← synthInstanceQ q(LinearOrderedField $α) + assertInstancesCommute + have z : Q(ℤ) := mkRawIntLit ⌈q⌉ + letI : $z =Q (-(-$n / $d)) := ⟨⟩ + return .isInt q(inferInstance) z ⌈q⌉ q(isInt_intCeil_ofIsRat _ $n $d $h) + | _, _, _ => failure + end NormNum end Rat diff --git a/MathlibTest/norm_num_ext.lean b/MathlibTest/norm_num_ext.lean index de2675f268b14..0e78dab0f080c 100644 --- a/MathlibTest/norm_num_ext.lean +++ b/MathlibTest/norm_num_ext.lean @@ -10,6 +10,7 @@ import Mathlib.Tactic.NormNum.DivMod import Mathlib.Tactic.NormNum.NatFib import Mathlib.Tactic.NormNum.NatSqrt import Mathlib.Tactic.NormNum.Prime +import Mathlib.Data.Rat.Floor import Mathlib.Tactic.NormNum.LegendreSymbol import Mathlib.Tactic.NormNum.Pow @@ -400,6 +401,11 @@ example : ⌊(2 : R)⌋ = 2 := by norm_num example : ⌊(15 / 16 : K)⌋ + 1 = 1 := by norm_num example : ⌊(-15 / 16 : K)⌋ + 1 = 0 := by norm_num +example : ⌈(-1 : R)⌉ = -1 := by norm_num +example : ⌈(2 : R)⌉ = 2 := by norm_num +example : ⌈(15 / 16 : K)⌉ + 1 = 2 := by norm_num +example : ⌈(-15 / 16 : K)⌉ + 1 = 1 := by norm_num + end floor section jacobi From 64f356003c7a06b16cb2767706a3b53a4303feaf Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Tue, 10 Dec 2024 02:59:32 +0000 Subject: [PATCH 621/829] chore: address and delete many porting notes across the library (#19685) --- Archive/Imo/Imo1972Q5.lean | 3 +- Archive/Imo/Imo2013Q1.lean | 1 - Archive/MiuLanguage/DecisionNec.lean | 8 ++--- Archive/MiuLanguage/DecisionSuf.lean | 6 ++-- Mathlib/Algebra/AddTorsor.lean | 4 --- Mathlib/Algebra/BigOperators/Finsupp.lean | 6 ---- .../Algebra/BigOperators/Group/Finset.lean | 3 +- .../Algebra/BigOperators/Group/Multiset.lean | 4 +-- Mathlib/Algebra/Category/Grp/Colimits.lean | 3 +- .../Grp/EquivalenceGroupAddGroup.lean | 6 ---- Mathlib/Algebra/CubicDiscriminant.lean | 8 ++--- Mathlib/Algebra/DirectLimit.lean | 9 ++--- Mathlib/Algebra/GeomSum.lean | 1 - Mathlib/Algebra/Lie/Submodule.lean | 6 ---- Mathlib/Algebra/Quandle.lean | 1 - Mathlib/Algebra/Quaternion.lean | 11 +++---- Mathlib/Analysis/Normed/Lp/LpEquiv.lean | 6 ---- .../Bicategory/Functor/Prelax.lean | 4 --- Mathlib/CategoryTheory/Limits/HasLimits.lean | 9 ++--- Mathlib/Data/List/DropRight.lean | 11 +------ Mathlib/Data/Set/Basic.lean | 1 - Mathlib/GroupTheory/GroupAction/Hom.lean | 33 +------------------ .../LinearAlgebra/FiniteDimensional/Defs.lean | 11 ------- .../Matrix/SesquilinearForm.lean | 1 - .../MeasureTheory/Measure/FiniteMeasure.lean | 5 +-- Mathlib/RingTheory/WittVector/IsPoly.lean | 15 --------- Mathlib/Topology/Gluing.lean | 25 ++++---------- Mathlib/Topology/UniformSpace/Completion.lean | 19 ++--------- 28 files changed, 31 insertions(+), 189 deletions(-) diff --git a/Archive/Imo/Imo1972Q5.lean b/Archive/Imo/Imo1972Q5.lean index ba2e4d8c6a0a0..75a0cc13edea4 100644 --- a/Archive/Imo/Imo1972Q5.lean +++ b/Archive/Imo/Imo1972Q5.lean @@ -80,11 +80,10 @@ This is a more concise version of the proof proposed by Ruben Van de Velde. -/ theorem imo1972_q5' (f g : ℝ → ℝ) (hf1 : ∀ x, ∀ y, f (x + y) + f (x - y) = 2 * f x * g y) (hf2 : BddAbove (Set.range fun x => ‖f x‖)) (hf3 : ∃ x, f x ≠ 0) (y : ℝ) : ‖g y‖ ≤ 1 := by - -- Porting note: moved `by_contra!` up to avoid a bug - by_contra! H obtain ⟨x, hx⟩ := hf3 set k := ⨆ x, ‖f x‖ have h : ∀ x, ‖f x‖ ≤ k := le_ciSup hf2 + by_contra! H 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) diff --git a/Archive/Imo/Imo2013Q1.lean b/Archive/Imo/Imo2013Q1.lean index 295886d7e18ea..6cd9697d983b5 100644 --- a/Archive/Imo/Imo2013Q1.lean +++ b/Archive/Imo/Imo2013Q1.lean @@ -27,7 +27,6 @@ We prove a slightly more general version where k does not need to be strictly po namespace Imo2013Q1 --- Porting note: simplified proof using `positivity` theorem arith_lemma (k n : ℕ) : 0 < 2 * n + 2 ^ k.succ := by positivity theorem prod_lemma (m : ℕ → ℕ+) (k : ℕ) (nm : ℕ+) : diff --git a/Archive/MiuLanguage/DecisionNec.lean b/Archive/MiuLanguage/DecisionNec.lean index 3bd06ca5463c6..e9c3d760485de 100644 --- a/Archive/MiuLanguage/DecisionNec.lean +++ b/Archive/MiuLanguage/DecisionNec.lean @@ -154,9 +154,7 @@ theorem goodm_of_rule3 (as bs : Miustr) (h₁ : Derivable (as ++ ↑[I, I, I] ++ exact mhead · contrapose! nmtail rcases exists_cons_of_ne_nil k with ⟨x, xs, rfl⟩ - -- Porting note: `simp_rw [cons_append]` didn't work - rw [cons_append] at nmtail; rw [cons_append, cons_append] - dsimp only [tail] at nmtail ⊢ + simp_rw [cons_append] at nmtail ⊢ simpa using nmtail /-! @@ -174,9 +172,7 @@ theorem goodm_of_rule4 (as bs : Miustr) (h₁ : Derivable (as ++ ↑[U, U] ++ bs exact mhead · contrapose! nmtail rcases exists_cons_of_ne_nil k with ⟨x, xs, rfl⟩ - -- Porting note: `simp_rw [cons_append]` didn't work - rw [cons_append] at nmtail; rw [cons_append, cons_append] - dsimp only [tail] at nmtail ⊢ + simp_rw [cons_append] at nmtail ⊢ simpa using nmtail /-- Any derivable string must begin with `M` and have no `M` in its tail. diff --git a/Archive/MiuLanguage/DecisionSuf.lean b/Archive/MiuLanguage/DecisionSuf.lean index 2f1f8bbf87750..4313004779cb8 100644 --- a/Archive/MiuLanguage/DecisionSuf.lean +++ b/Archive/MiuLanguage/DecisionSuf.lean @@ -308,10 +308,8 @@ theorem ind_hyp_suf (k : ℕ) (ys : Miustr) (hu : count U ys = succ k) (hdec : D rcases eq_append_cons_U_of_count_U_pos hu with ⟨as, bs, rfl⟩ use as, bs refine ⟨rfl, ?_, ?_, ?_⟩ - · -- Porting note: `simp_rw [count_append]` didn't work - rw [count_append] at hu - simp_rw [count_cons, beq_self_eq_true, if_true, add_succ, beq_iff_eq, reduceCtorEq, reduceIte, - add_zero, succ_inj'] at hu + · simp_rw [count_append, count_cons, beq_self_eq_true, if_true, add_succ, beq_iff_eq, + reduceCtorEq, reduceIte, add_zero, succ_inj'] at hu rwa [count_append, count_append] · apply And.intro rfl rw [cons_append, cons_append] diff --git a/Mathlib/Algebra/AddTorsor.lean b/Mathlib/Algebra/AddTorsor.lean index cbcdf82f7c7b7..c853185785a74 100644 --- a/Mathlib/Algebra/AddTorsor.lean +++ b/Mathlib/Algebra/AddTorsor.lean @@ -87,7 +87,6 @@ theorem vadd_vsub (g : G) (p : P) : (g +ᵥ p) -ᵥ p = g := /-- If the same point added to two group elements produces equal results, those group elements are equal. -/ theorem vadd_right_cancel {g₁ g₂ : G} (p : P) (h : g₁ +ᵥ p = g₂ +ᵥ p) : g₁ = g₂ := by --- Porting note: vadd_vsub g₁ → vadd_vsub g₁ p rw [← vadd_vsub g₁ p, h, vadd_vsub] @[simp] @@ -364,8 +363,6 @@ def constVAddHom : Multiplicative G →* Equiv.Perm P where variable {P} --- Porting note: Previous code was: --- open _Root_.Function open Function /-- Point reflection in `x` as a permutation. -/ @@ -412,7 +409,6 @@ theorem pointReflection_fixed_iff_of_injective_two_nsmul {x y : P} (h : Injectiv @[deprecated (since := "2024-11-18")] alias pointReflection_fixed_iff_of_injective_bit0 := pointReflection_fixed_iff_of_injective_two_nsmul --- Porting note: need this to calm down CI theorem injective_pointReflection_left_of_injective_two_nsmul {G P : Type*} [AddCommGroup G] [AddTorsor G P] (h : Injective (2 • · : G → G)) (y : P) : Injective fun x : P => pointReflection x y := diff --git a/Mathlib/Algebra/BigOperators/Finsupp.lean b/Mathlib/Algebra/BigOperators/Finsupp.lean index b992f81eade0b..fa5e090a17199 100644 --- a/Mathlib/Algebra/BigOperators/Finsupp.lean +++ b/Mathlib/Algebra/BigOperators/Finsupp.lean @@ -242,12 +242,10 @@ theorem sum_apply [Zero M] [AddCommMonoid N] {f : α →₀ M} {g : α → M → (f.sum g) a₂ = f.sum fun a₁ b => g a₁ b a₂ := finset_sum_apply _ _ _ --- Porting note: inserted ⇑ on the rhs @[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 @[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 _ _ @@ -381,14 +379,12 @@ theorem univ_sum_single [Fintype α] [AddCommMonoid M] (f : α →₀ M) : @[simp] theorem univ_sum_single_apply [AddCommMonoid M] [Fintype α] (i : α) (m : M) : ∑ j : α, single i m j = m := by - -- Porting note: rewrite due to leaky classical in lean3 classical rw [single, coe_mk, Finset.sum_pi_single'] simp @[simp] theorem univ_sum_single_apply' [AddCommMonoid M] [Fintype α] (i : α) (m : M) : ∑ j : α, single j m i = m := by - -- Porting note: rewrite due to leaky classical in lean3 simp_rw [single, coe_mk] classical rw [Finset.sum_pi_single] simp @@ -454,7 +450,6 @@ theorem support_sum_eq_biUnion {α : Type*} {ι : Type*} {M : Type*} [DecidableE (h : ∀ i₁ i₂, i₁ ≠ i₂ → Disjoint (g i₁).support (g i₂).support) : (∑ i ∈ s, g i).support = s.biUnion fun i => (g i).support := by classical - -- Porting note: apply Finset.induction_on s was not working; refine does. refine Finset.induction_on s ?_ ?_ · simp · intro i s hi @@ -578,7 +573,6 @@ end namespace Nat --- Porting note: Needed to replace pow with (· ^ ·) /-- If `0 : ℕ` is not in the support of `f : ℕ →₀ ℕ` then `0 < ∏ x ∈ f.support, x ^ (f x)`. -/ theorem prod_pow_pos_of_zero_not_mem_support {f : ℕ →₀ ℕ} (nhf : 0 ∉ f.support) : 0 < f.prod (· ^ ·) := diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index e6b48d65deaa9..88216bdc9801a 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -1216,8 +1216,7 @@ lemma prod_mulIndicator_subset_of_eq_one [One α] (f : ι → α) (g : ι → α ∏ i ∈ t, g i (mulIndicator ↑s f i) = ∏ i ∈ s, g i (f i) := by calc _ = ∏ i ∈ s, g i (mulIndicator ↑s f i) := by rw [prod_subset h fun i _ hn ↦ by simp [hn, hg]] - -- Porting note: This did not use to need the implicit argument - _ = _ := prod_congr rfl fun i hi ↦ congr_arg _ <| mulIndicator_of_mem (α := ι) hi f + _ = _ := prod_congr rfl fun i hi ↦ congr_arg _ <| mulIndicator_of_mem hi f /-- Taking the product of an indicator function over a possibly larger finset is the same as taking the original function over the original finset. -/ diff --git a/Mathlib/Algebra/BigOperators/Group/Multiset.lean b/Mathlib/Algebra/BigOperators/Group/Multiset.lean index a0bb921c056e3..a93afea870a07 100644 --- a/Mathlib/Algebra/BigOperators/Group/Multiset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Multiset.lean @@ -187,7 +187,6 @@ theorem prod_induction (p : α → Prop) (s : Multiset α) (p_mul : ∀ a b, p a @[to_additive] theorem prod_induction_nonempty (p : α → Prop) (p_mul : ∀ a b, p a → p b → p (a * b)) (hs : s ≠ ∅) (p_s : ∀ a ∈ s, p a) : p s.prod := by - -- Porting note: used to be `refine' Multiset.induction _ _` induction s using Multiset.induction_on with | empty => simp at hs | cons a s hsa => @@ -263,8 +262,7 @@ theorem prod_map_inv' (m : Multiset α) : (m.map Inv.inv).prod = m.prod⁻¹ := @[to_additive (attr := simp)] theorem prod_map_inv : (m.map fun i => (f i)⁻¹).prod = (m.map f).prod⁻¹ := by - -- Porting note: used `convert` - simp_rw [← (m.map f).prod_map_inv', map_map, Function.comp_apply] + rw [← (m.map f).prod_map_inv', map_map, Function.comp_def] @[to_additive (attr := simp)] theorem prod_map_div : (m.map fun i => f i / g i).prod = (m.map f).prod / (m.map g).prod := diff --git a/Mathlib/Algebra/Category/Grp/Colimits.lean b/Mathlib/Algebra/Category/Grp/Colimits.lean index ffd0441416766..1164b843e4f0b 100644 --- a/Mathlib/Algebra/Category/Grp/Colimits.lean +++ b/Mathlib/Algebra/Category/Grp/Colimits.lean @@ -209,8 +209,7 @@ def descFun (s : Cocone F) : ColimitType.{w} F → s.pt := by def descMorphism (s : Cocone F) : colimit.{w} F ⟶ s.pt where toFun := descFun F s map_zero' := rfl - -- Porting note: in `mathlib3`, nothing needs to be done after `induction` - map_add' x y := Quot.induction_on₂ x y fun _ _ => by dsimp; rw [← quot_add F]; rfl + map_add' x y := Quot.induction_on₂ x y fun _ _ ↦ rfl /-- Evidence that the proposed colimit is the colimit. -/ def colimitCoconeIsColimit : IsColimit (colimitCocone.{w} F) where diff --git a/Mathlib/Algebra/Category/Grp/EquivalenceGroupAddGroup.lean b/Mathlib/Algebra/Category/Grp/EquivalenceGroupAddGroup.lean index 1a42f23763056..892709d29b993 100644 --- a/Mathlib/Algebra/Category/Grp/EquivalenceGroupAddGroup.lean +++ b/Mathlib/Algebra/Category/Grp/EquivalenceGroupAddGroup.lean @@ -20,12 +20,6 @@ open CategoryTheory namespace Grp --- Porting note: Lean cannot find these now -private instance (X : Grp) : MulOneClass X.α := X.str.toMulOneClass -private instance (X : CommGrp) : MulOneClass X.α := X.str.toMulOneClass -private instance (X : AddGrp) : AddZeroClass X.α := X.str.toAddZeroClass -private instance (X : AddCommGrp) : AddZeroClass X.α := X.str.toAddZeroClass - /-- The functor `Grp ⥤ AddGrp` by sending `X ↦ Additive X` and `f ↦ f`. -/ @[simps] diff --git a/Mathlib/Algebra/CubicDiscriminant.lean b/Mathlib/Algebra/CubicDiscriminant.lean index 4104d2add1285..04cd0d3ed3474 100644 --- a/Mathlib/Algebra/CubicDiscriminant.lean +++ b/Mathlib/Algebra/CubicDiscriminant.lean @@ -468,12 +468,8 @@ def disc {R : Type*} [Ring R] (P : Cubic R) : R := theorem disc_eq_prod_three_roots (ha : P.a ≠ 0) (h3 : (map φ P).roots = {x, y, z}) : φ P.disc = (φ P.a * φ P.a * (x - y) * (x - z) * (y - z)) ^ 2 := by - simp only [disc, RingHom.map_add, RingHom.map_sub, RingHom.map_mul, map_pow] - -- Porting note: Replaced `simp only [RingHom.map_one, map_bit0, map_bit1]` with f4, f18, f27 - have f4 : φ 4 = 4 := map_natCast φ 4 - have f18 : φ 18 = 18 := map_natCast φ 18 - have f27 : φ 27 = 27 := map_natCast φ 27 - rw [f4, f18, f27, b_eq_three_roots ha h3, c_eq_three_roots ha h3, d_eq_three_roots ha h3] + simp only [disc, RingHom.map_add, RingHom.map_sub, RingHom.map_mul, map_pow, map_ofNat] + rw [b_eq_three_roots ha h3, c_eq_three_roots ha h3, d_eq_three_roots ha h3] ring1 theorem disc_ne_zero_iff_roots_ne (ha : P.a ≠ 0) (h3 : (map φ P).roots = {x, y, z}) : diff --git a/Mathlib/Algebra/DirectLimit.lean b/Mathlib/Algebra/DirectLimit.lean index af9c8c6a6689c..66ae30fb5b8fd 100644 --- a/Mathlib/Algebra/DirectLimit.lean +++ b/Mathlib/Algebra/DirectLimit.lean @@ -315,11 +315,7 @@ theorem of.zero_exact [IsDirected ι (· ≤ ·)] {i x} (H : of R ι G f i x = 0 if hx0 : x = 0 then ⟨i, le_rfl, by simp [hx0]⟩ else have hij : i ≤ j := hj _ <| by simp [DirectSum.apply_eq_component, hx0] - ⟨j, hij, by - -- Porting note: this had been - -- simpa [totalize_of_le hij] using hxj - simp only [DirectSum.toModule_lof] at hxj - rwa [totalize_of_le hij] at hxj⟩ + ⟨j, hij, by simpa [totalize, dif_pos hij, totalize_of_le hij] using hxj⟩ end DirectLimit @@ -731,8 +727,7 @@ theorem of.zero_exact_aux [Nonempty ι] [IsDirected ι (· ≤ ·)] {x : FreeCom ⟨k, ↑s ∪ t, this, isSupported_mul (isSupported_upwards hxs Set.subset_union_left) (isSupported_upwards hyt Set.subset_union_right), fun [_] => ?_⟩ - -- Porting note: RingHom.map_mul was `(restriction _).map_mul` - classical rw [RingHom.map_mul, (FreeCommRing.lift _).map_mul, ← + classical rw [(restriction _).map_mul, (FreeCommRing.lift _).map_mul, ← of.zero_exact_aux2 G f' hyt hj this hjk Set.subset_union_right, iht, (f' j k hjk).map_zero, mul_zero] diff --git a/Mathlib/Algebra/GeomSum.lean b/Mathlib/Algebra/GeomSum.lean index 7e6e9ce83070c..2a0df0564ca2b 100644 --- a/Mathlib/Algebra/GeomSum.lean +++ b/Mathlib/Algebra/GeomSum.lean @@ -346,7 +346,6 @@ protected theorem Commute.geom_sum₂_Ico_mul [Ring α] {x y : α} (h : Commute have hp := Commute.pow_pow (Commute.op h.symm) (n - 1 - k) k simpa [Commute, SemiconjBy] using hp simp only [this] - -- Porting note: gives deterministic timeout without this intermediate `have` convert (Commute.op h).mul_geom_sum₂_Ico hmn theorem geom_sum_Ico_mul [Ring α] (x : α) {m n : ℕ} (hmn : m ≤ n) : diff --git a/Mathlib/Algebra/Lie/Submodule.lean b/Mathlib/Algebra/Lie/Submodule.lean index 32a16eb6c461c..3408e81f11011 100644 --- a/Mathlib/Algebra/Lie/Submodule.lean +++ b/Mathlib/Algebra/Lie/Submodule.lean @@ -1209,15 +1209,9 @@ lemma incl_injective (I : LieIdeal R L) : Function.Injective I.incl := @[simp] theorem comap_incl_self : comap I.incl I = ⊤ := by ext; simp - -- porting note: `ext; simp` works also in mathlib3, though the proof used to be - -- rw [← LieSubmodule.coe_toSubmodule_eq_iff, LieSubmodule.top_coeSubmodule, - -- LieIdeal.comap_coeSubmodule, LieIdeal.incl_coe, Submodule.comap_subtype_self] @[simp] theorem ker_incl : I.incl.ker = ⊥ := by ext; simp - -- porting note: `ext; simp` works also in mathlib3, though the proof used to be - -- rw [← LieSubmodule.coe_toSubmodule_eq_iff, I.incl.ker_coeSubmodule, - -- LieSubmodule.bot_coeSubmodule, incl_coe, Submodule.ker_subtype] @[simp] theorem incl_idealRange : I.incl.idealRange = I := by diff --git a/Mathlib/Algebra/Quandle.lean b/Mathlib/Algebra/Quandle.lean index d02c6b1972aa1..fb77c232c8337 100644 --- a/Mathlib/Algebra/Quandle.lean +++ b/Mathlib/Algebra/Quandle.lean @@ -417,7 +417,6 @@ def Conj.map {G : Type*} {H : Type*} [Group G] [Group H] (f : G →* H) : Conj G Used for Fox n-colorings of knots. -/ --- Porting note: Removed nolint def Dihedral (n : ℕ) := ZMod n diff --git a/Mathlib/Algebra/Quaternion.lean b/Mathlib/Algebra/Quaternion.lean index e32a58000bb51..9a6904b709e90 100644 --- a/Mathlib/Algebra/Quaternion.lean +++ b/Mathlib/Algebra/Quaternion.lean @@ -148,7 +148,7 @@ theorem coe_injective : Function.Injective (coe : R → ℍ[R,c₁,c₂]) := fun theorem coe_inj {x y : R} : (x : ℍ[R,c₁,c₂]) = y ↔ x = y := coe_injective.eq_iff --- Porting note: removed `simps`, added simp lemmas manually +-- Porting note: removed `simps`, added simp lemmas manually. Should adjust `simps` to name properly instance : Zero ℍ[R,c₁,c₂] := ⟨⟨0, 0, 0, 0⟩⟩ @[simp] theorem zero_re : (0 : ℍ[R,c₁,c₂]).re = 0 := rfl @@ -169,7 +169,7 @@ instance : Inhabited ℍ[R,c₁,c₂] := ⟨0⟩ section One variable [One R] --- Porting note: removed `simps`, added simp lemmas manually +-- Porting note: removed `simps`, added simp lemmas manually. Should adjust `simps` to name properly instance : One ℍ[R,c₁,c₂] := ⟨⟨1, 0, 0, 0⟩⟩ @[simp] theorem one_re : (1 : ℍ[R,c₁,c₂]).re = 1 := rfl @@ -190,7 +190,7 @@ end Zero section Add variable [Add R] --- Porting note: removed `simps`, added simp lemmas manually +-- Porting note: removed `simps`, added simp lemmas manually. Should adjust `simps` to name properly instance : Add ℍ[R,c₁,c₂] := ⟨fun a b => ⟨a.1 + b.1, a.2 + b.2, a.3 + b.3, a.4 + b.4⟩⟩ @@ -223,7 +223,7 @@ end AddZeroClass section Neg variable [Neg R] --- Porting note: removed `simps`, added simp lemmas manually +-- Porting note: removed `simps`, added simp lemmas manually. Should adjust `simps` to name properly instance : Neg ℍ[R,c₁,c₂] := ⟨fun a => ⟨-a.1, -a.2, -a.3, -a.4⟩⟩ @[simp] theorem neg_re : (-a).re = -a.re := rfl @@ -718,6 +718,7 @@ end QuaternionAlgebra def Quaternion (R : Type*) [One R] [Neg R] := QuaternionAlgebra R (-1) (-1) +@[inherit_doc] scoped[Quaternion] notation "ℍ[" R "]" => Quaternion R /-- The equivalence between the quaternions over `R` and `R × R × R × R`. -/ @@ -1224,8 +1225,6 @@ instance instGroupWithZero : GroupWithZero ℍ[R] := inv := Inv.inv inv_zero := by rw [instInv_inv, star_zero, smul_zero] mul_inv_cancel := fun a ha => by - -- Porting note: the aliased definition confuse TC search - letI : Semiring ℍ[R] := inferInstanceAs (Semiring ℍ[R,-1,-1]) rw [instInv_inv, Algebra.mul_smul_comm (normSq a)⁻¹ a (star a), self_mul_star, smul_coe, inv_mul_cancel₀ (normSq_ne_zero.2 ha), coe_one] } diff --git a/Mathlib/Analysis/Normed/Lp/LpEquiv.lean b/Mathlib/Analysis/Normed/Lp/LpEquiv.lean index 6635a5d5fb355..d77bc7fea6e3e 100644 --- a/Mathlib/Analysis/Normed/Lp/LpEquiv.lean +++ b/Mathlib/Analysis/Normed/Lp/LpEquiv.lean @@ -91,9 +91,6 @@ section Equivₗᵢ variable [Fintype α] (𝕜 : Type*) [NontriviallyNormedField 𝕜] [∀ i, NormedSpace 𝕜 (E i)] variable (E) -/- porting note: Lean is unable to work with `lpPiLpₗᵢ` if `E` is implicit without -annotating with `(E := E)` everywhere, so we just make it explicit. This file has no -dependencies. -/ /-- The canonical `LinearIsometryEquiv` between `lp E p` and `PiLp p E` when `E : α → Type u` with `[Fintype α]` and `[Fact (1 ≤ p)]`. -/ @@ -145,9 +142,6 @@ theorem coe_addEquiv_lpBCF_symm (f : α →ᵇ E) : (AddEquiv.lpBCF.symm f : α rfl variable (E) -/- porting note: Lean is unable to work with `lpPiLpₗᵢ` if `E` is implicit without -annotating with `(E := E)` everywhere, so we just make it explicit. This file has no -dependencies. -/ /-- The canonical map between `lp (fun _ : α ↦ E) ∞` and `α →ᵇ E` as a `LinearIsometryEquiv`. -/ noncomputable def lpBCFₗᵢ : lp (fun _ : α ↦ E) ∞ ≃ₗᵢ[𝕜] α →ᵇ E := diff --git a/Mathlib/CategoryTheory/Bicategory/Functor/Prelax.lean b/Mathlib/CategoryTheory/Bicategory/Functor/Prelax.lean index f7bb524702e5c..a5bf157bf3564 100644 --- a/Mathlib/CategoryTheory/Bicategory/Functor/Prelax.lean +++ b/Mathlib/CategoryTheory/Bicategory/Functor/Prelax.lean @@ -76,10 +76,6 @@ def mkOfHomPrefunctors (F : B → C) (F' : (a : B) → (b : B) → Prefunctor (a map {a b} := (F' a b).obj map₂ {a b} := (F' a b).map --- Porting note: deleted syntactic tautologies `toPrefunctor_eq_coe : F.toPrefunctor = F` --- and `to_prefunctor_obj : (F : Prefunctor B C).obj = F.obj` --- and `to_prefunctor_map` - /-- The identity lax prefunctor. -/ @[simps] def id (B : Type u₁) [Quiver.{v₁ + 1} B] [∀ a b : B, Quiver.{w₁ + 1} (a ⟶ b)] : diff --git a/Mathlib/CategoryTheory/Limits/HasLimits.lean b/Mathlib/CategoryTheory/Limits/HasLimits.lean index febff5598358d..744fd90f39333 100644 --- a/Mathlib/CategoryTheory/Limits/HasLimits.lean +++ b/Mathlib/CategoryTheory/Limits/HasLimits.lean @@ -992,11 +992,10 @@ end variable {G : J ⥤ C} (α : F ⟶ G) --- @[reassoc (attr := simp)] Porting note: now simp can prove these @[reassoc] theorem colimit.ι_map (j : J) : colimit.ι F j ≫ colim.map α = α.app j ≫ colimit.ι G j := by simp -@[simp] -- Porting note: proof adjusted to account for @[simps] on all fields of colim +@[simp] theorem colimit.map_desc (c : Cocone G) : colimMap α ≫ colimit.desc G c = colimit.desc F ((Cocones.precompose α).obj c) := by ext j @@ -1125,8 +1124,7 @@ def IsColimit.op {t : Cocone F} (P : IsColimit t) : IsLimit t.op where -/ def IsLimit.unop {t : Cone F.op} (P : IsLimit t) : IsColimit t.unop where desc s := (P.lift s.op).unop - fac s j := congrArg Quiver.Hom.unop (P.fac s.op (Opposite.op j)) - -- Porting note: thinks `op j` is `IsLimit.op j` + fac s j := congrArg Quiver.Hom.unop (P.fac s.op (.op j)) uniq s m w := by dsimp rw [← P.uniq s.op m.op] @@ -1140,8 +1138,7 @@ def IsLimit.unop {t : Cone F.op} (P : IsLimit t) : IsColimit t.unop where -/ def IsColimit.unop {t : Cocone F.op} (P : IsColimit t) : IsLimit t.unop where lift s := (P.desc s.op).unop - fac s j := congrArg Quiver.Hom.unop (P.fac s.op (Opposite.op j)) - -- Porting note: thinks `op j` is `IsLimit.op j` + fac s j := congrArg Quiver.Hom.unop (P.fac s.op (.op j)) uniq s m w := by dsimp rw [← P.uniq s.op m.op] diff --git a/Mathlib/Data/List/DropRight.lean b/Mathlib/Data/List/DropRight.lean index f55afc95cdd9f..ad7f1c36acb62 100644 --- a/Mathlib/Data/List/DropRight.lean +++ b/Mathlib/Data/List/DropRight.lean @@ -138,18 +138,9 @@ theorem dropWhile_eq_self_iff : dropWhile p l = l ↔ ∀ hl : 0 < l.length, ¬p rw [get] at this simp_rw [this] -/- porting note: This proof is longer than it used to be because `simp` refuses to rewrite - the `l ≠ []` condition if `hl` is not `intro`'d yet -/ @[simp] theorem rdropWhile_eq_self_iff : rdropWhile p l = l ↔ ∀ hl : l ≠ [], ¬p (l.getLast hl) := by - simp only [rdropWhile, reverse_eq_iff, dropWhile_eq_self_iff, getLast_eq_getElem] - refine ⟨fun h hl => ?_, fun h hl => ?_⟩ - · rw [← length_pos, ← length_reverse] at hl - have := h hl - rwa [get_reverse'] at this - · rw [length_reverse, length_pos] at hl - have := h hl - rwa [get_reverse'] + simp [rdropWhile, reverse_eq_iff, getLast_eq_getElem, Nat.pos_iff_ne_zero] variable (p) (l) diff --git a/Mathlib/Data/Set/Basic.lean b/Mathlib/Data/Set/Basic.lean index 8d28bac27f0b0..ae281b9e0c8a5 100644 --- a/Mathlib/Data/Set/Basic.lean +++ b/Mathlib/Data/Set/Basic.lean @@ -1007,7 +1007,6 @@ def subtypeInsertEquivOption /-! ### Lemmas about singletons -/ -/- porting note: instance was in core in Lean3 -/ instance : LawfulSingleton α (Set α) := ⟨fun x => Set.ext fun a => by simp only [mem_empty_iff_false, mem_insert_iff, or_false] diff --git a/Mathlib/GroupTheory/GroupAction/Hom.lean b/Mathlib/GroupTheory/GroupAction/Hom.lean index 2ce20e3e8812b..f10d91924ee1c 100644 --- a/Mathlib/GroupTheory/GroupAction/Hom.lean +++ b/Mathlib/GroupTheory/GroupAction/Hom.lean @@ -150,10 +150,6 @@ theorem map_smul {F M X Y : Type*} [SMul M X] [SMul M Y] (f : F) (c : M) (x : X) : f (c • x) = c • f x := map_smulₛₗ f c x --- attribute [simp] map_smulₛₗ - --- Porting note: removed has_coe_to_fun instance, coercions handled differently now - @[to_additive] instance : MulActionSemiHomClass (X →ₑ[φ] Y) φ X Y where map_smulₛₗ := MulActionHom.map_smul' @@ -166,8 +162,6 @@ namespace MulActionHom variable {φ X Y} variable {F : Type*} [FunLike F X Y] -/- porting note: inserted following def & instance for consistent coercion behaviour, -see also Algebra.Hom.Group -/ /-- Turn an element of a type `F` satisfying `MulActionSemiHomClass F φ X Y` into an actual `MulActionHom`. This is declared as the default coercion from `F` to `MulActionSemiHom φ X Y`. -/ @@ -417,14 +411,6 @@ abbrev DistribMulActionHomClass (F : Type*) (M : outParam Type*) namespace DistribMulActionHom -/- Porting note (https://github.com/leanprover-community/mathlib4/issues/11215): TODO decide whether the next two instances should be removed -Coercion is already handled by all the HomClass constructions I believe -/ --- instance coe : Coe (A →+[M] B) (A →+ B) := --- ⟨toAddMonoidHom⟩ - --- instance coe' : Coe (A →+[M] B) (A →[M] B) := --- ⟨toMulActionHom⟩ - instance : FunLike (A →ₑ+[φ] B) A B where coe m := m.toFun coe_injective' f g h := by @@ -439,8 +425,6 @@ instance : DistribMulActionSemiHomClass (A →ₑ+[φ] B) φ A B where variable {φ φ' A B B₁} variable {F : Type*} [FunLike F A B] -/- porting note: inserted following def & instance for consistent coercion behaviour, -see also Algebra.Hom.Group -/ /-- Turn an element of a type `F` satisfying `MulActionHomClass F M X Y` into an actual `MulActionHom`. This is declared as the default coercion from `F` to `MulActionHom M X Y`. -/ @[coe] @@ -518,9 +502,8 @@ theorem id_apply (x : A) : DistribMulActionHom.id M x = x := by variable {M C ψ χ} --- porting note: `simp` used to prove this, but now `change` is needed to push past the coercions instance : Zero (A →ₑ+[φ] B) := - ⟨{ (0 : A →+ B) with map_smul' := fun m _ => by change (0 : B) = (φ m) • (0 : B); rw [smul_zero]}⟩ + ⟨{ (0 : A →+ B) with map_smul' := fun m _ => by simp }⟩ instance : One (A →+[M] A) := ⟨DistribMulActionHom.id M⟩ @@ -659,18 +642,6 @@ abbrev MulSemiringActionHomClass namespace MulSemiringActionHom -/- Porting note (https://github.com/leanprover-community/mathlib4/issues/11215): TODO decide whether the next two instances should be removed -Coercion is already handled by all the HomClass constructions I believe -/ --- @[coe] --- instance coe : Coe (R →+*[M] S) (R →+* S) := --- ⟨toRingHom⟩ - --- @[coe] --- instance coe' : Coe (R →+*[M] S) (R →+[M] S) := --- ⟨toDistribMulActionHom⟩ - --- Porting note: removed has_coe_to_fun instance, coercions handled differently now - instance : FunLike (R →ₑ+*[φ] S) R S where coe m := m.toFun coe_injective' f g h := by @@ -687,8 +658,6 @@ instance : MulSemiringActionSemiHomClass (R →ₑ+*[φ] S) φ R S where variable {φ R S} variable {F : Type*} [FunLike F R S] -/- porting note: inserted following def & instance for consistent coercion behaviour, -see also Algebra.Hom.Group -/ /-- Turn an element of a type `F` satisfying `MulSemiringActionHomClass F M R S` into an actual `MulSemiringActionHom`. This is declared as the default coercion from `F` to `MulSemiringActionHom M X Y`. -/ diff --git a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean index ccf5f3bda0fb2..16ac3a373d3e2 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean @@ -758,17 +758,6 @@ open Module variable {F E : Type*} [Field F] [Ring E] [Algebra F E] -/- -porting note: -Some of the lemmas in this section can be made faster by adding these short-cut instances -```lean4 -instance (S : Subalgebra F E) : AddCommMonoid { x // x ∈ S } := inferInstance -instance (S : Subalgebra F E) : AddCommGroup { x // x ∈ S } := inferInstance -``` -However, this approach doesn't scale very well, so we should consider holding off on adding -them until we have no choice. --/ - /-- A `Subalgebra` is `FiniteDimensional` iff it is `FiniteDimensional` as a submodule. -/ theorem Subalgebra.finiteDimensional_toSubmodule {S : Subalgebra F E} : FiniteDimensional F (Subalgebra.toSubmodule S) ↔ FiniteDimensional F S := diff --git a/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean b/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean index d15dfc288b115..150f3e2f40e25 100644 --- a/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean +++ b/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean @@ -563,7 +563,6 @@ theorem Matrix.isAdjointPair_equiv (P : Matrix n n R) (h : IsUnit P) : dsimp only [Matrix.IsAdjointPair] simp only [Matrix.transpose_mul] simp only [← mul_assoc, P.transpose_nonsing_inv] - -- Porting note: the previous proof used `conv` and was causing timeouts, so we use `convert` convert this using 2 · rw [mul_assoc, mul_assoc, ← mul_assoc J] rfl diff --git a/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean b/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean index 1b17d6219d497..babc5786d359e 100644 --- a/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean @@ -411,10 +411,7 @@ theorem continuous_testAgainstNN_eval (f : Ω →ᵇ ℝ≥0) : Continuous fun μ : FiniteMeasure Ω ↦ μ.testAgainstNN f := by show Continuous ((fun φ : WeakDual ℝ≥0 (Ω →ᵇ ℝ≥0) ↦ φ f) ∘ toWeakDualBCNN) refine Continuous.comp ?_ (toWeakDualBCNN_continuous (Ω := Ω)) - exact WeakBilin.eval_continuous (𝕜 := ℝ≥0) (E := (Ω →ᵇ ℝ≥0) →L[ℝ≥0] ℝ≥0) _ _ - /- porting note: without explicitly providing `𝕜` and `E` TC synthesis times - out trying to find `Module ℝ≥0 ((Ω →ᵇ ℝ≥0) →L[ℝ≥0] ℝ≥0)`, but it can find it with enough time: - `set_option synthInstance.maxHeartbeats 47000` was sufficient. -/ + exact WeakBilin.eval_continuous _ _ /-- The total mass of a finite measure depends continuously on the measure. -/ theorem continuous_mass : Continuous fun μ : FiniteMeasure Ω ↦ μ.mass := by diff --git a/Mathlib/RingTheory/WittVector/IsPoly.lean b/Mathlib/RingTheory/WittVector/IsPoly.lean index eed73bc325570..d364756920a22 100644 --- a/Mathlib/RingTheory/WittVector/IsPoly.lean +++ b/Mathlib/RingTheory/WittVector/IsPoly.lean @@ -337,20 +337,6 @@ namespace IsPoly₂ instance [Fact p.Prime] : Inhabited (IsPoly₂ p (fun _ _ => (· + ·))) := ⟨addIsPoly₂⟩ --- Porting note: maybe just drop this now that it works by `inferInstance` -/-- The composition of a binary polynomial function - with a unary polynomial function in the first argument is polynomial. -/ -theorem compLeft {g f} [IsPoly₂ p g] [IsPoly p f] : - IsPoly₂ p fun _R _Rcr x y => g (f x) y := - inferInstance - --- Porting note: maybe just drop this now that it works by `inferInstance` -/-- The composition of a binary polynomial function - with a unary polynomial function in the second argument is polynomial. -/ -theorem compRight {g f} [IsPoly₂ p g] [IsPoly p f] : - IsPoly₂ p fun _R _Rcr x y => g x (f y) := - inferInstance - theorem ext [Fact p.Prime] {f g} (hf : IsPoly₂ p f) (hg : IsPoly₂ p g) (h : ∀ (R : Type u) [_Rcr : CommRing R] (x y : 𝕎 R) (n : ℕ), ghostComponent n (f x y) = ghostComponent n (g x y)) : @@ -360,7 +346,6 @@ theorem ext [Fact p.Prime] {f g} (hf : IsPoly₂ p f) (hg : IsPoly₂ p g) intros ext n rw [hf, hg, poly_eq_of_wittPolynomial_bind_eq' p φ ψ] - -- porting note: `clear x y` does not work, since `x, y` are now hygienic intro k apply MvPolynomial.funext intro x diff --git a/Mathlib/Topology/Gluing.lean b/Mathlib/Topology/Gluing.lean index 6c4f75721330d..9af7fc85de01f 100644 --- a/Mathlib/Topology/Gluing.lean +++ b/Mathlib/Topology/Gluing.lean @@ -100,11 +100,9 @@ theorem isOpen_iff (U : Set 𝖣.glued) : IsOpen U ↔ ∀ i, IsOpen (𝖣.ι i delta CategoryTheory.GlueData.ι simp_rw [← Multicoequalizer.ι_sigmaπ 𝖣.diagram] rw [← (homeoOfIso (Multicoequalizer.isoCoequalizer 𝖣.diagram).symm).isOpen_preimage] - rw [coequalizer_isOpen_iff] + rw [coequalizer_isOpen_iff, colimit_isOpen_iff.{u}] dsimp only [GlueData.diagram_l, GlueData.diagram_left, GlueData.diagram_r, GlueData.diagram_right, parallelPair_obj_one] - rw [colimit_isOpen_iff.{_,u}] -- Porting note: changed `.{u}` to `.{_,u}`. fun fact: the proof - -- breaks down if this `rw` is merged with the `rw` above. constructor · intro h j; exact h ⟨j⟩ · intro h j; cases j; apply h @@ -274,17 +272,14 @@ theorem preimage_image_eq_image' (i j : D.J) (U : Set (𝖣.U i)) : -- Porting note: `show` was not needed, since `rw [← Set.image_image]` worked. show (fun x => ((forget TopCat).map _ ((forget TopCat).map _ x))) '' _ = _ rw [← Set.image_image] - -- Porting note: `congr 1` was here, instead of `congr_arg`, however, it did nothing. - refine congr_arg ?_ ?_ + congr! 1 rw [← Set.eq_preimage_iff_image_eq, Set.preimage_preimage] · change _ = (D.t i j ≫ D.t j i ≫ _) ⁻¹' _ rw [𝖣.t_inv_assoc] rw [← isIso_iff_bijective] apply (forget TopCat).map_isIso --- Porting note: the goal was simply `IsOpen (𝖣.ι i '' U)`. --- I had to manually add the explicit type ascription. -theorem open_image_open (i : D.J) (U : Opens (𝖣.U i)) : IsOpen (𝖣.ι i '' (U : Set (D.U i))) := by +theorem open_image_open (i : D.J) (U : Opens (𝖣.U i)) : IsOpen (𝖣.ι i '' U) := by rw [isOpen_iff] intro j rw [preimage_image_eq_image] @@ -393,15 +388,8 @@ def ofOpenSubsets : TopCat.GlueData.{u} := { J U := fun i => (Opens.toTopCat <| TopCat.of α).obj (U i) V := fun _ j => (Opens.map <| Opens.inclusion' _).obj (U j) - t := fun i j => ⟨fun x => ⟨⟨x.1.1, x.2⟩, x.1.2⟩, by - -- Porting note: was `continuity`, see https://github.com/leanprover-community/mathlib4/issues/5030 - refine Continuous.subtype_mk ?_ ?_ - refine Continuous.subtype_mk ?_ ?_ - continuity⟩ - V_id := fun i => by - ext - -- Porting note: no longer needed `cases U i`! - simp + t := fun i j => ⟨fun x => ⟨⟨x.1.1, x.2⟩, x.1.2⟩, by fun_prop⟩ + V_id := fun i => by ext; simp t_id := fun i => by ext; rfl t_inter := fun _ _ _ _ hx => hx cocycle := fun _ _ _ _ _ => rfl } @@ -445,8 +433,7 @@ theorem fromOpenSubsetsGlue_isOpenMap : IsOpenMap (fromOpenSubsetsGlue U) := by apply (Opens.isOpenEmbedding (X := TopCat.of α) (U i)).isOpenMap convert hs i using 1 erw [← ι_fromOpenSubsetsGlue, coe_comp, Set.preimage_comp] - -- porting note: `congr 1` did nothing, so I replaced it with `apply congr_arg` - apply congr_arg + congr! 1 exact Set.preimage_image_eq _ (fromOpenSubsetsGlue_injective U) · refine ⟨Set.mem_image_of_mem _ hx, ?_⟩ rw [ι_fromOpenSubsetsGlue_apply] diff --git a/Mathlib/Topology/UniformSpace/Completion.lean b/Mathlib/Topology/UniformSpace/Completion.lean index b10f71d844724..bfea3bac43b00 100644 --- a/Mathlib/Topology/UniformSpace/Completion.lean +++ b/Mathlib/Topology/UniformSpace/Completion.lean @@ -204,9 +204,6 @@ theorem nonempty_cauchyFilter_iff : Nonempty (CauchyFilter α) ↔ Nonempty α : section --- Porting note: I commented this --- set_option eqn_compiler.zeta true - instance : CompleteSpace (CauchyFilter α) := completeSpace_extension isUniformInducing_pureCauchy denseRange_pureCauchy fun f hf => let f' : CauchyFilter α := ⟨f, hf⟩ @@ -313,9 +310,7 @@ instance completeSpace : CompleteSpace (Completion α) := instance t0Space : T0Space (Completion α) := SeparationQuotient.instT0Space -/-- The map from a uniform space to its completion. - -porting note: this was added to create a target for the `@[coe]` attribute. -/ +/-- The map from a uniform space to its completion. -/ @[coe] def coe' : α → Completion α := SeparationQuotient.mk ∘ pureCauchy /-- Automatic coercion from `α` to its completion. Not always injective. -/ @@ -460,8 +455,6 @@ theorem continuous_extension : Continuous (Completion.extension f) := end CompleteSpace -/- porting note: removed `@[simp]` because this lemma doesn't even trigger on itself in Lean 3 or -Lean 4 unless the user manually supplies the `hf` argument, so it is useless as a `simp` lemma. -/ theorem extension_coe [T0Space β] (hf : UniformContinuous f) (a : α) : (Completion.extension f) a = f a := cPkg.extend_coe hf a @@ -495,8 +488,6 @@ theorem uniformContinuous_map : UniformContinuous (Completion.map f) := theorem continuous_map : Continuous (Completion.map f) := cPkg.continuous_map cPkg f -/- porting note: removed `@[simp]` because this lemma doesn't even trigger on itself in Lean 3 or -Lean 4 unless the user manually supplies the `hf` argument, so it is useless as a `simp` lemma. -/ theorem map_coe (hf : UniformContinuous f) (a : α) : (Completion.map f) a = f a := cPkg.map_coe cPkg hf a @@ -512,11 +503,7 @@ theorem extension_map [CompleteSpace γ] [T0Space γ] {f : β → γ} {g : α (hf : UniformContinuous f) (hg : UniformContinuous g) : Completion.extension f ∘ Completion.map g = Completion.extension (f ∘ g) := Completion.ext (continuous_extension.comp continuous_map) continuous_extension <| by - intro a - -- Porting note: this is not provable by simp [hf, hg, hf.comp hg, map_coe, extension_coe], - -- but should be? - rw [extension_coe (hf.comp hg), Function.comp_apply, map_coe hg, extension_coe hf, - Function.comp_apply] + simp [hf, hg, hf.comp hg, map_coe, extension_coe] theorem map_comp {g : β → γ} {f : α → β} (hg : UniformContinuous g) (hf : UniformContinuous f) : Completion.map g ∘ Completion.map f = Completion.map (g ∘ f) := @@ -568,8 +555,6 @@ section T0Space variable [T0Space γ] {f} -/- porting note: removed `@[simp]` because this lemma doesn't even trigger on itself in Lean 3 or -Lean 4 unless the user manually supplies the `hf` argument, so it is useless as a `simp` lemma. -/ theorem extension₂_coe_coe (hf : UniformContinuous₂ f) (a : α) (b : β) : Completion.extension₂ f a b = f a b := cPkg.extension₂_coe_coe cPkg hf a b From 078c1ad7f056414c7fc02601f03963312ed99624 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 10 Dec 2024 02:59:33 +0000 Subject: [PATCH 622/829] feat: more `Int.ediv` lemmas (#19754) --- Mathlib/Algebra/Order/Group/Unbundled/Int.lean | 7 +++++++ Mathlib/Data/Int/Defs.lean | 5 +++++ Mathlib/Data/Int/Lemmas.lean | 12 ++++++++++++ 3 files changed, 24 insertions(+) diff --git a/Mathlib/Algebra/Order/Group/Unbundled/Int.lean b/Mathlib/Algebra/Order/Group/Unbundled/Int.lean index 96c6eed502379..2fad62292ee5a 100644 --- a/Mathlib/Algebra/Order/Group/Unbundled/Int.lean +++ b/Mathlib/Algebra/Order/Group/Unbundled/Int.lean @@ -46,6 +46,9 @@ theorem natAbs_abs (a : ℤ) : natAbs |a| = natAbs a := by rw [abs_eq_natAbs]; r theorem sign_mul_abs (a : ℤ) : sign a * |a| = a := by rw [abs_eq_natAbs, sign_mul_natAbs a] +theorem sign_mul_self_eq_abs (a : ℤ) : sign a * a = |a| := by + rw [abs_eq_natAbs, sign_mul_self_eq_natAbs] + lemma natAbs_le_self_sq (a : ℤ) : (Int.natAbs a : ℤ) ≤ a ^ 2 := by rw [← Int.natAbs_sq a, sq] norm_cast @@ -129,6 +132,10 @@ protected theorem sign_eq_ediv_abs (a : ℤ) : sign a = a / |a| := if az : a = 0 then by simp [az] else (Int.ediv_eq_of_eq_mul_left (mt abs_eq_zero.1 az) (sign_mul_abs _).symm).symm +protected theorem sign_eq_abs_ediv (a : ℤ) : sign a = |a| / a := + if az : a = 0 then by simp [az] + else (Int.ediv_eq_of_eq_mul_left az (sign_mul_self_eq_abs _).symm).symm + end Int section Group diff --git a/Mathlib/Data/Int/Defs.lean b/Mathlib/Data/Int/Defs.lean index 5fc0b54bd9a45..081c427a4817b 100644 --- a/Mathlib/Data/Int/Defs.lean +++ b/Mathlib/Data/Int/Defs.lean @@ -348,6 +348,11 @@ lemma natAbs_sq (x : ℤ) : (x.natAbs : ℤ) ^ 2 = x ^ 2 := by alias natAbs_pow_two := natAbs_sq +theorem sign_mul_self_eq_natAbs : ∀ a : Int, sign a * a = natAbs a + | 0 => rfl + | Nat.succ _ => Int.one_mul _ + | -[_+1] => (Int.neg_eq_neg_one_mul _).symm + /-! ### `/` -/ @[simp, norm_cast] lemma natCast_div (m n : ℕ) : ((m / n : ℕ) : ℤ) = m / n := rfl diff --git a/Mathlib/Data/Int/Lemmas.lean b/Mathlib/Data/Int/Lemmas.lean index 42c4167321de5..f98d34103dd0d 100644 --- a/Mathlib/Data/Int/Lemmas.lean +++ b/Mathlib/Data/Int/Lemmas.lean @@ -133,4 +133,16 @@ theorem div2_bit (b n) : div2 (bit b n) = n := by @[deprecated (since := "2024-04-02")] alias abs_coe_nat := abs_natCast @[deprecated (since := "2024-04-02")] alias coe_nat_nonpos_iff := natCast_nonpos_iff +/-- Like `Int.ediv_emod_unique`, but permitting negative `b`. -/ +theorem ediv_emod_unique' {a b r q : Int} (h : b ≠ 0) : + a / b = q ∧ a % b = r ↔ r + b * q = a ∧ 0 ≤ r ∧ r < |b| := by + constructor + · intro ⟨rfl, rfl⟩ + exact ⟨emod_add_ediv a b, emod_nonneg _ h, emod_lt _ h⟩ + · intro ⟨rfl, hz, hb⟩ + constructor + · rw [Int.add_mul_ediv_left r q h, ediv_eq_zero_of_lt_abs hz hb] + simp [Int.zero_add] + · rw [add_mul_emod_self_left, ← emod_abs, emod_eq_of_lt hz hb] + end Int From 23b5a98a5456a69277fe7193ae2c4b33bb652b93 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 10 Dec 2024 04:05:43 +0000 Subject: [PATCH 623/829] chore: split the nlinarith preprocessor for clarity (#19841) This makes it a little easier to trace the two halves separately, which is desirable since the second half has quadratic complexity. The behavior and performance should be unchanged. Co-authored-by: Eric Wieser --- Mathlib/Tactic/Linarith/Preprocessing.lean | 78 +++++++++++++--------- 1 file changed, 46 insertions(+), 32 deletions(-) diff --git a/Mathlib/Tactic/Linarith/Preprocessing.lean b/Mathlib/Tactic/Linarith/Preprocessing.lean index 348309a7774af..f714751349976 100644 --- a/Mathlib/Tactic/Linarith/Preprocessing.lean +++ b/Mathlib/Tactic/Linarith/Preprocessing.lean @@ -287,6 +287,50 @@ partial def findSquares (s : RBSet (Nat × Bool) lexOrd.compare) (e : Expr) : e.foldlM findSquares s | _ => e.foldlM findSquares s +/-- Get proofs of `-x^2 ≤ 0` and `-(x*x) ≤ 0`, when those terms appear in `ls` -/ +private def nlinarithGetSquareProofs (ls : List Expr) : MetaM (List Expr) := do + -- find the squares in `AtomM` to ensure deterministic behavior + let s ← AtomM.run .reducible do + let si ← ls.foldrM (fun h s' => do findSquares s' (← instantiateMVars (← inferType h))) + RBSet.empty + si.toList.mapM fun (i, is_sq) => return ((← get).atoms[i]!, is_sq) + let new_es ← s.filterMapM fun (e, is_sq) => + observing? <| mkAppM (if is_sq then ``sq_nonneg else ``mul_self_nonneg) #[e] + let new_es ← compWithZero.globalize.transform new_es + trace[linarith] "nlinarith preprocessing found squares" + trace[linarith] "{s}" + linarithTraceProofs "so we added proofs" new_es + return new_es + +/-- +Get proofs for products of inequalities from `ls`. + +Note that the length of the resulting list is proportional to `ls.length^2`, which can make a large +amount of work for the linarith oracle. +-/ +private def nlinarithGetProductsProofs (ls : List Expr) : MetaM (List Expr) := do + let with_comps ← ls.mapM (fun e => do + let tp ← inferType e + try + let ⟨ine, _⟩ ← parseCompAndExpr tp + pure (ine, e) + catch _ => pure (Ineq.lt, e)) + let products ← with_comps.mapDiagM fun (⟨posa, a⟩ : Ineq × Expr) ⟨posb, b⟩ => + try + (some <$> match posa, posb with + | Ineq.eq, _ => mkAppM ``zero_mul_eq #[a, b] + | _, Ineq.eq => mkAppM ``mul_zero_eq #[a, b] + | Ineq.lt, Ineq.lt => mkAppM ``mul_pos_of_neg_of_neg #[a, b] + | Ineq.lt, Ineq.le => do + let a ← mkAppM ``le_of_lt #[a] + mkAppM ``mul_nonneg_of_nonpos_of_nonpos #[a, b] + | Ineq.le, Ineq.lt => do + let b ← mkAppM ``le_of_lt #[b] + mkAppM ``mul_nonneg_of_nonpos_of_nonpos #[a, b] + | Ineq.le, Ineq.le => mkAppM ``mul_nonneg_of_nonpos_of_nonpos #[a, b]) + catch _ => pure none + compWithZero.globalize.transform products.reduceOption + /-- `nlinarithExtras` is the preprocessor corresponding to the `nlinarith` tactic. @@ -299,38 +343,8 @@ This preprocessor is typically run last, after all inputs have been canonized. def nlinarithExtras : GlobalPreprocessor where name := "nonlinear arithmetic extras" transform ls := do - -- find the squares in `AtomM` to ensure deterministic behavior - let s ← AtomM.run .reducible do - let si ← ls.foldrM (fun h s' => do findSquares s' (← instantiateMVars (← inferType h))) - RBSet.empty - si.toList.mapM fun (i, is_sq) => return ((← get).atoms[i]!, is_sq) - let new_es ← s.filterMapM fun (e, is_sq) => - observing? <| mkAppM (if is_sq then ``sq_nonneg else ``mul_self_nonneg) #[e] - let new_es ← compWithZero.globalize.transform new_es - trace[linarith] "nlinarith preprocessing found squares" - trace[linarith] "{s}" - linarithTraceProofs "so we added proofs" new_es - let with_comps ← (new_es ++ ls).mapM (fun e => do - let tp ← inferType e - try - let ⟨ine, _⟩ ← parseCompAndExpr tp - pure (ine, e) - catch _ => pure (Ineq.lt, e)) - let products ← with_comps.mapDiagM fun (⟨posa, a⟩ : Ineq × Expr) ⟨posb, b⟩ => - try - (some <$> match posa, posb with - | Ineq.eq, _ => mkAppM ``zero_mul_eq #[a, b] - | _, Ineq.eq => mkAppM ``mul_zero_eq #[a, b] - | Ineq.lt, Ineq.lt => mkAppM ``mul_pos_of_neg_of_neg #[a, b] - | Ineq.lt, Ineq.le => do - let a ← mkAppM ``le_of_lt #[a] - mkAppM ``mul_nonneg_of_nonpos_of_nonpos #[a, b] - | Ineq.le, Ineq.lt => do - let b ← mkAppM ``le_of_lt #[b] - mkAppM ``mul_nonneg_of_nonpos_of_nonpos #[a, b] - | Ineq.le, Ineq.le => mkAppM ``mul_nonneg_of_nonpos_of_nonpos #[a, b]) - catch _ => pure none - let products ← compWithZero.globalize.transform products.reduceOption + let new_es ← nlinarithGetSquareProofs ls + let products ← nlinarithGetProductsProofs (new_es ++ ls) return (new_es ++ ls ++ products) end nlinarith From c90bcde4894d81cbda7f29fdccd41532ef24c2ac Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Tue, 10 Dec 2024 05:49:29 +0000 Subject: [PATCH 624/829] feat: add some nontriviality results on tensor product (#19832) - `TensorProduct.nontrivial_of_linearMap_injective_of_flat_(left|right)`: if there is an injective linear map from the base ring to a module, and another module is nontrivial and flat, then their tensor product is nontrivial. - `Algebra.TensorProduct.nontrivial_of_algebraMap_injective_of_flat_(left|right)`: if there is an algebra whose algebra map is injective, and there is another nontrivial flat algebra, then their tensor product is nontrivial. - `Algebra.TensorProduct.nontrivial_of_algebraMap_injective_of_isDomain`: tensor product of two domains is nontrivial, provided that two algebra maps are injective. --- Mathlib.lean | 1 + Mathlib/RingTheory/Flat/Basic.lean | 50 +++++++++++++++++-- .../RingTheory/TensorProduct/Nontrivial.lean | 41 +++++++++++++++ 3 files changed, 89 insertions(+), 3 deletions(-) create mode 100644 Mathlib/RingTheory/TensorProduct/Nontrivial.lean diff --git a/Mathlib.lean b/Mathlib.lean index 07fe8c7ea85df..1c288bfd59f6c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4571,6 +4571,7 @@ import Mathlib.RingTheory.TensorProduct.Basic import Mathlib.RingTheory.TensorProduct.Finite import Mathlib.RingTheory.TensorProduct.Free import Mathlib.RingTheory.TensorProduct.MvPolynomial +import Mathlib.RingTheory.TensorProduct.Nontrivial import Mathlib.RingTheory.TensorProduct.Pi import Mathlib.RingTheory.Trace.Basic import Mathlib.RingTheory.Trace.Defs diff --git a/Mathlib/RingTheory/Flat/Basic.lean b/Mathlib/RingTheory/Flat/Basic.lean index a36eb5785365c..18985920c6140 100644 --- a/Mathlib/RingTheory/Flat/Basic.lean +++ b/Mathlib/RingTheory/Flat/Basic.lean @@ -56,11 +56,13 @@ See . universe v' u v w +open TensorProduct + namespace Module open Function (Surjective) -open LinearMap Submodule TensorProduct DirectSum +open LinearMap Submodule DirectSum variable (R : Type u) (M : Type v) [CommRing R] [AddCommGroup M] [Module R M] @@ -426,8 +428,6 @@ end Module section Injective -open scoped TensorProduct - variable {R S A B : Type*} [CommRing R] [Ring A] [Algebra R A] [Ring B] [Algebra R B] [CommSemiring S] [Algebra S A] [SMulCommClass R S A] @@ -448,3 +448,47 @@ theorem includeRight_injective [Module.Flat R B] (ha : Function.Injective (algeb end Algebra.TensorProduct end Injective + +section Nontrivial + +variable (R : Type*) [CommRing R] + +namespace TensorProduct + +variable (M N : Type*) [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] + +/-- If `M`, `N` are `R`-modules, there exists an injective `R`-linear map from `R` to `N`, +and `M` is a nontrivial flat `R`-module, then `M ⊗[R] N` is nontrivial. -/ +theorem nontrivial_of_linearMap_injective_of_flat_left (f : R →ₗ[R] N) (h : Function.Injective f) + [Module.Flat R M] [Nontrivial M] : Nontrivial (M ⊗[R] N) := + Module.Flat.lTensor_preserves_injective_linearMap (M := M) f h |>.comp + (TensorProduct.rid R M).symm.injective |>.nontrivial + +/-- If `M`, `N` are `R`-modules, there exists an injective `R`-linear map from `R` to `M`, +and `N` is a nontrivial flat `R`-module, then `M ⊗[R] N` is nontrivial. -/ +theorem nontrivial_of_linearMap_injective_of_flat_right (f : R →ₗ[R] M) (h : Function.Injective f) + [Module.Flat R N] [Nontrivial N] : Nontrivial (M ⊗[R] N) := + Module.Flat.rTensor_preserves_injective_linearMap (M := N) f h |>.comp + (TensorProduct.lid R N).symm.injective |>.nontrivial + +end TensorProduct + +namespace Algebra.TensorProduct + +variable (A B : Type*) [CommRing A] [CommRing B] [Algebra R A] [Algebra R B] + +/-- If `A`, `B` are `R`-algebras, `R` injects into `B`, +and `A` is a nontrivial flat `R`-algebra, then `A ⊗[R] B` is nontrivial. -/ +theorem nontrivial_of_algebraMap_injective_of_flat_left (h : Function.Injective (algebraMap R B)) + [Module.Flat R A] [Nontrivial A] : Nontrivial (A ⊗[R] B) := + TensorProduct.nontrivial_of_linearMap_injective_of_flat_left R A B (Algebra.linearMap R B) h + +/-- If `A`, `B` are `R`-algebras, `R` injects into `A`, +and `B` is a nontrivial flat `R`-algebra, then `A ⊗[R] B` is nontrivial. -/ +theorem nontrivial_of_algebraMap_injective_of_flat_right (h : Function.Injective (algebraMap R A)) + [Module.Flat R B] [Nontrivial B] : Nontrivial (A ⊗[R] B) := + TensorProduct.nontrivial_of_linearMap_injective_of_flat_right R A B (Algebra.linearMap R A) h + +end Algebra.TensorProduct + +end Nontrivial diff --git a/Mathlib/RingTheory/TensorProduct/Nontrivial.lean b/Mathlib/RingTheory/TensorProduct/Nontrivial.lean new file mode 100644 index 0000000000000..ee304a244cbe6 --- /dev/null +++ b/Mathlib/RingTheory/TensorProduct/Nontrivial.lean @@ -0,0 +1,41 @@ +/- +Copyright (c) 2024 Jz Pan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jz Pan +-/ +import Mathlib.LinearAlgebra.Basis.VectorSpace +import Mathlib.RingTheory.Flat.FaithfullyFlat +import Mathlib.RingTheory.Localization.FractionRing + +/-! + +# Nontriviality of tensor product of algebras + +This file contains some more results on nontriviality of tensor product of algebras. + +-/ + +open TensorProduct + +namespace Algebra.TensorProduct + +/-- If `A`, `B` are `R`-algebras, `R` injects into `A` and `B`, and `A` and `B` are domains +(which implies `R` is also a domain), then `A ⊗[R] B` is nontrivial. -/ +theorem nontrivial_of_algebraMap_injective_of_isDomain + (R A B : Type*) [CommRing R] [CommRing A] [CommRing B] [Algebra R A] [Algebra R B] + (ha : Function.Injective (algebraMap R A)) (hb : Function.Injective (algebraMap R B)) + [IsDomain A] [IsDomain B] : Nontrivial (A ⊗[R] B) := by + haveI := ha.isDomain _ + let FR := FractionRing R + let FA := FractionRing A + let FB := FractionRing B + let fa : FR →ₐ[R] FA := IsFractionRing.liftAlgHom (g := Algebra.ofId R FA) + ((IsFractionRing.injective A FA).comp ha) + let fb : FR →ₐ[R] FB := IsFractionRing.liftAlgHom (g := Algebra.ofId R FB) + ((IsFractionRing.injective B FB).comp hb) + algebraize_only [fa.toRingHom, fb.toRingHom] + exact Algebra.TensorProduct.mapOfCompatibleSMul FR R FA FB |>.comp + (Algebra.TensorProduct.map (IsScalarTower.toAlgHom R A FA) (IsScalarTower.toAlgHom R B FB)) + |>.toRingHom.domain_nontrivial + +end Algebra.TensorProduct From a456587dace5491b17f36eb824cdc717a41c3e5a Mon Sep 17 00:00:00 2001 From: David Renshaw Date: Tue, 10 Dec 2024 06:05:21 +0000 Subject: [PATCH 625/829] chore(CategoryTheory/MorphismProperty/IsInvertedBy): simplify proof of isoClosure_iff (#19854) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * `have : f = e.inv.left ≫ f' ≫ e.hom.right` can be directly proved via [`Arrow.iso_w'`](https://github.com/leanprover-community/mathlib4/blob/ebd71de74db22aaa0abbfab6799a22bcecda3476/Mathlib/CategoryTheory/Comma/Arrow.lean#L156). This removes an `erw`. * Once the `have` is shortened like that, we might as well inline it at its usage in the subsequent `simp`. This change was suggested by [tryAtEachStep](https://github.com/dwrensha/tryAtEachStep). --- Mathlib/CategoryTheory/MorphismProperty/IsInvertedBy.lean | 5 +---- 1 file changed, 1 insertion(+), 4 deletions(-) diff --git a/Mathlib/CategoryTheory/MorphismProperty/IsInvertedBy.lean b/Mathlib/CategoryTheory/MorphismProperty/IsInvertedBy.lean index 44711f734e411..16ed860f8bcb4 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/IsInvertedBy.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/IsInvertedBy.lean @@ -126,10 +126,7 @@ lemma IsInvertedBy.isoClosure_iff (W : MorphismProperty C) (F : C ⥤ D) : · intro h X Y f hf exact h _ (W.le_isoClosure _ hf) · intro h X Y f ⟨X', Y', f', hf', ⟨e⟩⟩ - have : f = e.inv.left ≫ f' ≫ e.hom.right := by - erw [← e.hom.w, ← Arrow.comp_left_assoc, e.inv_hom_id, Category.id_comp] - rfl - simp only [this, F.map_comp] + simp only [Arrow.iso_w' e, F.map_comp] have := h _ hf' infer_instance From 619e12537cb09ff0bfab5ed5c49ccb4d1b875623 Mon Sep 17 00:00:00 2001 From: Scott Carnahan <128885296+ScottCarnahan@users.noreply.github.com> Date: Tue, 10 Dec 2024 06:18:38 +0000 Subject: [PATCH 626/829] refactor(LinearAlgebra/RootSystem): Weyl group (#19744) This PR redefines the Weyl group of a root pairing as the subgroup of automorphisms generated by reflection morphisms. The new definition is isomorphic to the old one, by faithfulness of the weight space representation, but we also get a natural coweight space representation and permutation representation on roots. Co-authored-by: Oliver Nash --- Mathlib.lean | 1 + Mathlib/LinearAlgebra/RootSystem/Defs.lean | 78 +---------- Mathlib/LinearAlgebra/RootSystem/Hom.lean | 76 ++++++++-- .../LinearAlgebra/RootSystem/WeylGroup.lean | 131 ++++++++++++++++++ 4 files changed, 202 insertions(+), 84 deletions(-) create mode 100644 Mathlib/LinearAlgebra/RootSystem/WeylGroup.lean diff --git a/Mathlib.lean b/Mathlib.lean index 1c288bfd59f6c..59289e49c814b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3469,6 +3469,7 @@ import Mathlib.LinearAlgebra.RootSystem.Hom import Mathlib.LinearAlgebra.RootSystem.OfBilinear import Mathlib.LinearAlgebra.RootSystem.RootPairingCat import Mathlib.LinearAlgebra.RootSystem.RootPositive +import Mathlib.LinearAlgebra.RootSystem.WeylGroup import Mathlib.LinearAlgebra.SModEq import Mathlib.LinearAlgebra.Semisimple import Mathlib.LinearAlgebra.SesquilinearForm diff --git a/Mathlib/LinearAlgebra/RootSystem/Defs.lean b/Mathlib/LinearAlgebra/RootSystem/Defs.lean index da02ee54a1daf..39c9941d22e0f 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Defs.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Defs.lean @@ -413,14 +413,6 @@ lemma corootSpan_dualAnnihilator_map_eq : (span R (range P.coroot')).dualCoannihilator := P.flip.rootSpan_dualAnnihilator_map_eq -/-- The `Weyl group` of a root pairing is the group of automorphisms of the weight space generated -by reflections in roots. -/ -def weylGroup : Subgroup (M ≃ₗ[R] M) := - Subgroup.closure (range P.reflection) - -lemma reflection_mem_weylGroup : P.reflection i ∈ P.weylGroup := - Subgroup.subset_closure <| mem_range_self i - lemma mem_range_root_of_mem_range_reflection_of_mem_range_root {r : M ≃ₗ[R] M} {α : M} (hr : r ∈ range P.reflection) (hα : α ∈ range P.root) : r • α ∈ range P.root := by @@ -435,71 +427,6 @@ lemma mem_range_coroot_of_mem_range_coreflection_of_mem_range_coroot obtain ⟨j, rfl⟩ := hα exact ⟨P.reflection_perm i j, P.coroot_reflection_perm i j⟩ -lemma exists_root_eq_smul_of_mem_weylGroup {w : M ≃ₗ[R] M} (hw : w ∈ P.weylGroup) (i : ι) : - ∃ j, P.root j = w • P.root i := - Subgroup.smul_mem_of_mem_closure_of_mem (by simp) - (fun _ h _ ↦ P.mem_range_root_of_mem_range_reflection_of_mem_range_root h) hw (mem_range_self i) - -/-- The permutation representation of the Weyl group induced by `reflection_perm`. -/ -def weylGroupToPerm : P.weylGroup →* Equiv.Perm ι where - toFun w := - { toFun := fun i => (P.exists_root_eq_smul_of_mem_weylGroup w.2 i).choose - invFun := fun i => (P.exists_root_eq_smul_of_mem_weylGroup w⁻¹.2 i).choose - left_inv := fun i => by - obtain ⟨w, hw⟩ := w - apply P.root.injective - rw [(P.exists_root_eq_smul_of_mem_weylGroup ((Subgroup.inv_mem_iff P.weylGroup).mpr hw) - ((P.exists_root_eq_smul_of_mem_weylGroup hw i).choose)).choose_spec, - (P.exists_root_eq_smul_of_mem_weylGroup hw i).choose_spec, inv_smul_smul] - right_inv := fun i => by - obtain ⟨w, hw⟩ := w - have hw' : w⁻¹ ∈ P.weylGroup := (Subgroup.inv_mem_iff P.weylGroup).mpr hw - apply P.root.injective - rw [(P.exists_root_eq_smul_of_mem_weylGroup hw - ((P.exists_root_eq_smul_of_mem_weylGroup hw' i).choose)).choose_spec, - (P.exists_root_eq_smul_of_mem_weylGroup hw' i).choose_spec, smul_inv_smul] } - map_one' := by ext; simp - map_mul' x y := by - obtain ⟨x, hx⟩ := x - obtain ⟨y, hy⟩ := y - ext i - apply P.root.injective - simp only [Equiv.coe_fn_mk, Equiv.Perm.coe_mul, comp_apply] - rw [(P.exists_root_eq_smul_of_mem_weylGroup (mul_mem hx hy) i).choose_spec, - (P.exists_root_eq_smul_of_mem_weylGroup hx - ((P.exists_root_eq_smul_of_mem_weylGroup hy i).choose)).choose_spec, - (P.exists_root_eq_smul_of_mem_weylGroup hy i).choose_spec, mul_smul] - -@[simp] -lemma weylGroupToPerm_apply_reflection : - P.weylGroupToPerm ⟨P.reflection i, P.reflection_mem_weylGroup i⟩ = P.reflection_perm i := by - ext j - apply P.root.injective - rw [weylGroupToPerm, MonoidHom.coe_mk, OneHom.coe_mk, Equiv.coe_fn_mk, root_reflection_perm, - (P.exists_root_eq_smul_of_mem_weylGroup (P.reflection_mem_weylGroup i) j).choose_spec, - LinearEquiv.smul_def] - -@[simp] -lemma range_weylGroupToPerm : - P.weylGroupToPerm.range = Subgroup.closure (range P.reflection_perm) := by - refine (Subgroup.closure_eq_of_le _ ?_ ?_).symm - · rintro - ⟨i, rfl⟩ - simpa only [← weylGroupToPerm_apply_reflection] using mem_range_self _ - · rintro - ⟨⟨w, hw⟩, rfl⟩ - induction hw using Subgroup.closure_induction'' with - | one => - change P.weylGroupToPerm 1 ∈ _ - simpa only [map_one] using Subgroup.one_mem _ - | mem w' hw' => - obtain ⟨i, rfl⟩ := hw' - simpa only [weylGroupToPerm_apply_reflection] using Subgroup.subset_closure (mem_range_self i) - | inv_mem w' hw' => - obtain ⟨i, rfl⟩ := hw' - simpa only [reflection_inv, weylGroupToPerm_apply_reflection] using - Subgroup.subset_closure (mem_range_self i) - | mul w₁ w₂ hw₁ hw₂ h₁ h₂ => - simpa only [← Submonoid.mk_mul_mk _ w₁ w₂ hw₁ hw₂, map_mul] using Subgroup.mul_mem _ h₁ h₂ - lemma pairing_smul_root_eq (k : ι) (hij : P.reflection_perm i = P.reflection_perm j) : P.pairing k i • P.root i = P.pairing k j • P.root j := by have h : P.reflection i (P.root k) = P.reflection j (P.root k) := by @@ -573,9 +500,12 @@ lemma exists_int_eq_coxeterWeight [P.IsCrystallographic] (i j : ι) : /-- Two roots are orthogonal when they are fixed by each others' reflections. -/ def IsOrthogonal : Prop := pairing P i j = 0 ∧ pairing P j i = 0 -lemma IsOrthogonal.symm : IsOrthogonal P i j ↔ IsOrthogonal P j i := by +lemma isOrthogonal_symm : IsOrthogonal P i j ↔ IsOrthogonal P j i := by simp only [IsOrthogonal, and_comm] +lemma IsOrthogonal.symm (h : IsOrthogonal P i j) : IsOrthogonal P j i := + ⟨h.2, h.1⟩ + lemma isOrthogonal_comm (h : IsOrthogonal P i j) : Commute (P.reflection i) (P.reflection j) := by rw [commute_iff_eq] ext v diff --git a/Mathlib/LinearAlgebra/RootSystem/Hom.lean b/Mathlib/LinearAlgebra/RootSystem/Hom.lean index 08f552e72ccc3..1a8565d1abfea 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Hom.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Hom.lean @@ -235,6 +235,12 @@ lemma coweightHom_injective (P : RootPairing ι R M N) : Injective (coweightHom rw [Equiv.symm_apply_apply] at this rw [this, Equiv.apply_symm_apply] +/-- The permutation representation of the endomorphism monoid on the root index set -/ +def indexHom (P : RootPairing ι R M N) : End P →* (ι ≃ ι) where + toFun f := Hom.indexEquiv f + map_one' := by ext; simp + map_mul' x y := by ext; simp + end Hom variable {ι₂ M₂ N₂ : Type*} @@ -349,7 +355,7 @@ lemma comp_assoc {ι₁ M₁ N₁ ι₂ M₂ N₂ ι₃ M₃ N₃ : Type*} [AddC comp (comp h g) f = comp h (comp g f) := by ext <;> simp -/-- The endomorphism monoid of a root pairing. -/ +/-- Equivalences form a monoid. -/ instance (P : RootPairing ι R M N) : Monoid (RootPairing.Equiv P P) where mul := comp mul_assoc := comp_assoc @@ -454,7 +460,7 @@ lemma inv_indexEquiv {ι₂ M₂ N₂ : Type*} [AddCommGroup M₂] [Module R M (f : RootPairing.Equiv P Q) : (symm P Q f).indexEquiv = (Hom.indexEquiv f.toHom).symm := rfl -/-- The endomorphism monoid of a root pairing. -/ +/-- Equivalences form a group. -/ instance (P : RootPairing ι R M N) : Group (RootPairing.Equiv P P) where mul := comp mul_assoc := comp_assoc @@ -528,21 +534,26 @@ def weightHom (P : RootPairing ι R M N) : Aut P →* (M ≃ₗ[R] M) where map_mul' x y := by ext; simp lemma weightHom_toLinearMap {P : RootPairing ι R M N} (g : Aut P) : - ((weightHom P) g).toLinearMap = (Hom.weightHom P) g.toHom := + (weightHom P g).toLinearMap = Hom.weightHom P g.toHom := rfl lemma weightHom_injective (P : RootPairing ι R M N) : Injective (Equiv.weightHom P) := by refine Injective.of_comp (f := LinearEquiv.toLinearMap) fun g g' hgg' => ?_ - let h : ((weightHom P) g).toLinearMap = ((weightHom P) g').toLinearMap := hgg' --`have` gets lint + let h : (weightHom P g).toLinearMap = (weightHom P g').toLinearMap := hgg' --`have` gets lint rw [weightHom_toLinearMap, weightHom_toLinearMap] at h suffices h' : g.toHom = g'.toHom by exact Equiv.ext hgg' (congrArg Hom.coweightMap h') (congrArg Hom.indexEquiv h') exact Hom.weightHom_injective P hgg' +@[simp] +lemma weightEquiv_inv {P : RootPairing ι R M N} (g : Aut P) : + weightEquiv P P g⁻¹ = (weightEquiv P P g)⁻¹ := + LinearEquiv.toLinearMap_inj.mp rfl + /-- The coweight space representation of automorphisms -/ @[simps] def coweightHom (P : RootPairing ι R M N) : Aut P →* (N ≃ₗ[R] N)ᵐᵒᵖ where - toFun g := MulOpposite.op ((coweightEquiv P P) g) + toFun g := MulOpposite.op (coweightEquiv P P g) map_one' := by simp only [MulOpposite.op_eq_one_iff] exact LinearEquiv.toLinearMap_inj.mp rfl @@ -551,14 +562,14 @@ def coweightHom (P : RootPairing ι R M N) : Aut P →* (N ≃ₗ[R] N)ᵐᵒᵖ exact fun x y ↦ rfl lemma coweightHom_toLinearMap {P : RootPairing ι R M N} (g : Aut P) : - (MulOpposite.unop ((coweightHom P) g)).toLinearMap = - MulOpposite.unop ((Hom.coweightHom P) g.toHom) := + (MulOpposite.unop (coweightHom P g)).toLinearMap = + MulOpposite.unop (Hom.coweightHom P g.toHom) := rfl lemma coweightHom_injective (P : RootPairing ι R M N) : Injective (Equiv.coweightHom P) := by refine Injective.of_comp (f := fun a => MulOpposite.op a) fun g g' hgg' => ?_ - have h : (MulOpposite.unop ((coweightHom P) g)).toLinearMap = - (MulOpposite.unop ((coweightHom P) g')).toLinearMap := by + have h : (MulOpposite.unop (coweightHom P g)).toLinearMap = + (MulOpposite.unop (coweightHom P g')).toLinearMap := by simp_all rw [coweightHom_toLinearMap, coweightHom_toLinearMap] at h suffices h' : g.toHom = g'.toHom by @@ -566,6 +577,27 @@ lemma coweightHom_injective (P : RootPairing ι R M N) : Injective (Equiv.coweig apply Hom.coweightHom_injective P exact MulOpposite.unop_inj.mp h +lemma coweightHom_op {P : RootPairing ι R M N} (g : Aut P) : + MulOpposite.unop (coweightHom P g) = coweightEquiv P P g := + rfl + +@[simp] +lemma coweightEquiv_inv {P : RootPairing ι R M N} (g : Aut P) : + coweightEquiv P P g⁻¹ = (coweightEquiv P P g)⁻¹ := + LinearEquiv.toLinearMap_inj.mp rfl + +/-- The permutation representation of the automorphism group on the root index set -/ +@[simps] +def indexHom (P : RootPairing ι R M N) : Aut P →* (ι ≃ ι) where + toFun g := g.toHom.indexEquiv + map_one' := by ext; simp + map_mul' x y := by ext; simp + +@[simp] +lemma indexEquiv_inv {P : RootPairing ι R M N} (g : Aut P) : + (g⁻¹).toHom.indexEquiv = (indexHom P g)⁻¹ := + rfl + /-- The automorphism of a root pairing given by a reflection. -/ def reflection (P : RootPairing ι R M N) (i : ι) : Aut P where weightMap := P.reflection i @@ -601,7 +633,31 @@ lemma reflection_coweightEquiv (P : RootPairing ι R M N) (i : ι) : @[simp] lemma reflection_indexEquiv (P : RootPairing ι R M N) (i : ι) : - (reflection P i).indexEquiv = P.reflection_perm i := rfl + (reflection P i).indexEquiv = P.reflection_perm i := + rfl + +@[simp] +lemma reflection_inv (P : RootPairing ι R M N) (i : ι) : + (reflection P i)⁻¹ = (reflection P i) := by + refine Equiv.ext ?_ ?_ ?_ + · exact LinearMap.ext_iff.mpr (fun x => by simp [← weightEquiv_apply]) + · exact LinearMap.ext_iff.mpr (fun x => by simp [← coweightEquiv_apply]) + · exact _root_.Equiv.ext (fun j => by simp only [← indexHom_apply, map_inv]; simp) + +instance : MulAction P.Aut M where + smul w v := Equiv.weightHom P w v + one_smul _ := rfl + mul_smul _ _ _ := rfl + +instance : MulAction (P.Aut)ᵐᵒᵖ N where + smul w v := (MulOpposite.unop (Equiv.coweightHom P (MulOpposite.unop w))) v + one_smul _ := rfl + mul_smul _ _ _ := rfl + +instance : MulAction P.Aut ι where + smul w i := Equiv.indexHom P w i + one_smul _ := rfl + mul_smul _ _ _ := rfl end Equiv diff --git a/Mathlib/LinearAlgebra/RootSystem/WeylGroup.lean b/Mathlib/LinearAlgebra/RootSystem/WeylGroup.lean new file mode 100644 index 0000000000000..c89c7d4158b1a --- /dev/null +++ b/Mathlib/LinearAlgebra/RootSystem/WeylGroup.lean @@ -0,0 +1,131 @@ +/- +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.Hom + +/-! +# The Weyl group of a root pairing +This file defines the Weyl group of a root pairing as the subgroup of automorphisms generated by +reflection automorphisms. This deviates from the existing literature, which typically defines the +Weyl group as the subgroup of linear transformations of the weight space generated by linear +reflections. However, the automorphism group of a root pairing comes with a permutation +representation on the set indexing roots and faithful linear representations on the weight space and +coweight space. Thus, our formalism gives us an isomorphism to the traditional Weyl group together +with the natural dual representation generated by coreflections and the permutation representation +on roots. +## Main definitions + * `RootPairing.weylGroup` : The group of automorphisms generated by reflections. + * `RootPairing.weylGroupToPerm` : The permutation representation of the Weyl group on roots. +## Results + * `RootPairing.range_weylGroup_weightHom` : The image of the weight space representation is equal + to the subgroup generated by linear reflections. + * `RootPairing.range_weylGroup_coweightHom` : The image of the coweight space representation is + equal to the subgroup generated by linear coreflections. + * `RootPairing.range_weylGroupToPerm` : The image of the permutation representation is equal to the + subgroup generated by reflection permutations. +## TODO + * faithfulness of `weylGroupToPerm` when multiplication by 2 is injective on the weight space. +-/ + +open Set Function + +variable {ι R M N : Type*} + +noncomputable section + +namespace RootPairing + +variable [CommRing R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] + (P : RootPairing ι R M N) (i : ι) + +/-- The `Weyl group` of a root pairing is the group of automorphisms of the root pairing generated +by reflections. -/ +def weylGroup : Subgroup (Aut P) := + Subgroup.closure (range (Equiv.reflection P)) + +lemma reflection_mem_weylGroup : Equiv.reflection P i ∈ P.weylGroup := + Subgroup.subset_closure <| mem_range_self i + +lemma range_weylGroup_weightHom : + MonoidHom.range ((Equiv.weightHom P).restrict P.weylGroup) = + Subgroup.closure (range P.reflection) := by + refine (Subgroup.closure_eq_of_le _ ?_ ?_).symm + · rintro - ⟨i, rfl⟩ + simp only [MonoidHom.restrict_range, Subgroup.coe_map, Equiv.weightHom_apply, mem_image, + SetLike.mem_coe] + use Equiv.reflection P i + exact ⟨reflection_mem_weylGroup P i, Equiv.reflection_weightEquiv P i⟩ + · rintro fg ⟨⟨w, hw⟩, rfl⟩ + induction hw using Subgroup.closure_induction'' with + | one => + change ((Equiv.weightHom P).restrict P.weylGroup) 1 ∈ _ + simpa only [map_one] using Subgroup.one_mem _ + | mem w' hw' => + obtain ⟨i, rfl⟩ := hw' + simp only [MonoidHom.restrict_apply, Equiv.weightHom_apply, Equiv.reflection_weightEquiv] + simpa only [reflection_mem_weylGroup] using Subgroup.subset_closure (mem_range_self i) + | inv_mem w' hw' => + obtain ⟨i, rfl⟩ := hw' + simp only [Equiv.reflection_inv, MonoidHom.restrict_apply, Equiv.weightHom_apply, + Equiv.reflection_weightEquiv] + simpa only [reflection_mem_weylGroup] using Subgroup.subset_closure (mem_range_self i) + | mul w₁ w₂ hw₁ hw₂ h₁ h₂ => + simpa only [← Submonoid.mk_mul_mk _ w₁ w₂ hw₁ hw₂, map_mul] using Subgroup.mul_mem _ h₁ h₂ + +lemma range_weylGroup_coweightHom : + MonoidHom.range ((Equiv.coweightHom P).restrict P.weylGroup) = + Subgroup.closure (range (MulOpposite.op ∘ P.coreflection)) := by + refine (Subgroup.closure_eq_of_le _ ?_ ?_).symm + · rintro - ⟨i, rfl⟩ + simp only [MonoidHom.restrict_range, Subgroup.coe_map, Equiv.weightHom_apply, mem_image, + SetLike.mem_coe] + use Equiv.reflection P i + refine ⟨reflection_mem_weylGroup P i, by simp⟩ + · rintro fg ⟨⟨w, hw⟩, rfl⟩ + induction hw using Subgroup.closure_induction'' with + | one => + change ((Equiv.coweightHom P).restrict P.weylGroup) 1 ∈ _ + simpa only [map_one] using Subgroup.one_mem _ + | mem w' hw' => + obtain ⟨i, rfl⟩ := hw' + simp only [MonoidHom.restrict_apply, Equiv.coweightHom_apply, Equiv.reflection_coweightEquiv] + simpa only [reflection_mem_weylGroup] using Subgroup.subset_closure (mem_range_self i) + | inv_mem w' hw' => + obtain ⟨i, rfl⟩ := hw' + simp only [Equiv.reflection_inv, MonoidHom.restrict_apply, Equiv.coweightHom_apply, + Equiv.reflection_coweightEquiv] + simpa only [reflection_mem_weylGroup] using Subgroup.subset_closure (mem_range_self i) + | mul w₁ w₂ hw₁ hw₂ h₁ h₂ => + simpa only [← Submonoid.mk_mul_mk _ w₁ w₂ hw₁ hw₂, map_mul] using Subgroup.mul_mem _ h₁ h₂ + +/-- The permutation representation of the Weyl group induced by `reflection_perm`. -/ +abbrev weylGroupToPerm := (Equiv.indexHom P).restrict P.weylGroup + +lemma range_weylGroupToPerm : + P.weylGroupToPerm.range = Subgroup.closure (range P.reflection_perm) := by + refine (Subgroup.closure_eq_of_le _ ?_ ?_).symm + · rintro - ⟨i, rfl⟩ + simp only [MonoidHom.restrict_range, Subgroup.coe_map, Equiv.weightHom_apply, mem_image, + SetLike.mem_coe] + use Equiv.reflection P i + refine ⟨reflection_mem_weylGroup P i, by simp⟩ + · rintro fg ⟨⟨w, hw⟩, rfl⟩ + induction hw using Subgroup.closure_induction'' with + | one => + change ((Equiv.indexHom P).restrict P.weylGroup) 1 ∈ _ + simpa only [map_one] using Subgroup.one_mem _ + | mem w' hw' => + obtain ⟨i, rfl⟩ := hw' + simp only [MonoidHom.restrict_apply, Equiv.indexHom_apply, Equiv.reflection_indexEquiv] + simpa only [reflection_mem_weylGroup] using Subgroup.subset_closure (mem_range_self i) + | inv_mem w' hw' => + obtain ⟨i, rfl⟩ := hw' + simp only [Equiv.reflection_inv, MonoidHom.restrict_apply, Equiv.indexHom_apply, + Equiv.reflection_indexEquiv] + simpa only [reflection_mem_weylGroup] using Subgroup.subset_closure (mem_range_self i) + | mul w₁ w₂ hw₁ hw₂ h₁ h₂ => + simpa only [← Submonoid.mk_mul_mk _ w₁ w₂ hw₁ hw₂, map_mul] using Subgroup.mul_mem _ h₁ h₂ + +end RootPairing From 783026989dcdad61379854036e571fbe45ece043 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Tue, 10 Dec 2024 06:48:09 +0000 Subject: [PATCH 627/829] feat: more API for iteratedFDeriv (#19769) Notably, `iteratedFDeriv` commutes with composition with a translation --- Mathlib/Analysis/Analytic/ChangeOrigin.lean | 8 ++ Mathlib/Analysis/Calculus/ContDiff/Basic.lean | 4 +- Mathlib/Analysis/Calculus/ContDiff/Defs.lean | 34 ++++- .../Calculus/ContDiff/FTaylorSeries.lean | 82 +++++++++++- Mathlib/Analysis/Calculus/FDeriv/Add.lean | 118 ++++++++++++++++++ .../Analysis/Calculus/FDeriv/Analytic.lean | 35 ++++-- .../Calculus/FormalMultilinearSeries.lean | 10 ++ Mathlib/Data/Fintype/Card.lean | 22 +++- Mathlib/Topology/Homeomorph.lean | 3 + 9 files changed, 297 insertions(+), 19 deletions(-) diff --git a/Mathlib/Analysis/Analytic/ChangeOrigin.lean b/Mathlib/Analysis/Analytic/ChangeOrigin.lean index 45dce2a33326c..176f6f15e6afc 100644 --- a/Mathlib/Analysis/Analytic/ChangeOrigin.lean +++ b/Mathlib/Analysis/Analytic/ChangeOrigin.lean @@ -232,6 +232,14 @@ theorem radius_le_radius_derivSeries : p.radius ≤ p.derivSeries.radius := by apply mul_le_of_le_one_left (norm_nonneg _) exact ContinuousLinearMap.opNorm_le_bound _ zero_le_one (by simp) +theorem derivSeries_eq_zero {n : ℕ} (hp : p (n + 1) = 0) : p.derivSeries n = 0 := by + suffices p.changeOriginSeries 1 n = 0 by ext v; simp [derivSeries, this] + apply Finset.sum_eq_zero (fun s hs ↦ ?_) + ext v + have : p (1 + n) = 0 := p.congr_zero (by abel) hp + simp [changeOriginSeriesTerm, ContinuousMultilinearMap.curryFinFinset_apply, + ContinuousMultilinearMap.zero_apply, this] + end -- From this point on, assume that the space is complete, to make sure that series that converge diff --git a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean index 91cf96657a498..7b0fc7934c102 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean @@ -966,14 +966,14 @@ Warning: if you think you need this lemma, it is likely that you can simplify yo reformulating the lemma that you're applying next using the tips in Note [continuity lemma statement] -/ -theorem contDiff_prodAssoc : ContDiff 𝕜 ⊤ <| Equiv.prodAssoc E F G := +theorem contDiff_prodAssoc : ContDiff 𝕜 ω <| Equiv.prodAssoc E F G := (LinearIsometryEquiv.prodAssoc 𝕜 E F G).contDiff /-- The natural equivalence `E × (F × G) ≃ (E × F) × G` is smooth. Warning: see remarks attached to `contDiff_prodAssoc` -/ -theorem contDiff_prodAssoc_symm : ContDiff 𝕜 ⊤ <| (Equiv.prodAssoc E F G).symm := +theorem contDiff_prodAssoc_symm : ContDiff 𝕜 ω <| (Equiv.prodAssoc E F G).symm := (LinearIsometryEquiv.prodAssoc 𝕜 E F G).symm.contDiff /-! ### Bundled derivatives are smooth -/ diff --git a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean index 6f8cc0631548c..e772546295c8f 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean @@ -689,6 +689,11 @@ protected theorem ContDiffOn.ftaylorSeriesWithin exact (Hp.mono ho).eq_iteratedFDerivWithin_of_uniqueDiffOn le_rfl (hs.inter o_open) ⟨hy, yo⟩ exact ((Hp.mono ho).cont m le_rfl).congr fun y hy => (A y hy).symm +theorem iteratedFDerivWithin_subset {n : ℕ} (st : s ⊆ t) (hs : UniqueDiffOn 𝕜 s) + (ht : UniqueDiffOn 𝕜 t) (h : ContDiffOn 𝕜 n f t) (hx : x ∈ s) : + iteratedFDerivWithin 𝕜 n f s x = iteratedFDerivWithin 𝕜 n f t x := + (((h.ftaylorSeriesWithin ht).mono st).eq_iteratedFDerivWithin_of_uniqueDiffOn le_rfl hs hx).symm + /-- On a set with unique differentiability, an analytic function is automatically `C^ω`, as its successive derivatives are also analytic. This does not require completeness of the space. See also `AnalyticOn.contDiffOn_of_completeSpace`.-/ @@ -1015,6 +1020,20 @@ protected theorem ContDiffAt.eventually (h : ContDiffAt 𝕜 n f x) (h' : n ≠ ∀ᶠ y in 𝓝 x, ContDiffAt 𝕜 n f y := by simpa [nhdsWithin_univ] using ContDiffWithinAt.eventually h h' +theorem iteratedFDerivWithin_eq_iteratedFDeriv {n : ℕ} + (hs : UniqueDiffOn 𝕜 s) (h : ContDiffAt 𝕜 n f x) (hx : x ∈ s) : + iteratedFDerivWithin 𝕜 n f s x = iteratedFDeriv 𝕜 n f x := by + rw [← iteratedFDerivWithin_univ] + rcases h.contDiffOn' le_rfl (by simp) with ⟨u, u_open, xu, hu⟩ + rw [← iteratedFDerivWithin_inter_open u_open xu, + ← iteratedFDerivWithin_inter_open u_open xu (s := univ)] + apply iteratedFDerivWithin_subset + · exact inter_subset_inter_left _ (subset_univ _) + · exact hs.inter u_open + · apply uniqueDiffOn_univ.inter u_open + · simpa using hu + · exact ⟨hx, xu⟩ + /-! ### Smooth functions -/ variable (𝕜) in @@ -1144,15 +1163,22 @@ theorem contDiff_omega_iff_analyticOnNhd : /-! ### Iterated derivative -/ +/-- When a function is `C^n`, it admits `ftaylorSeries 𝕜 f` as a Taylor series up +to order `n` in `s`. -/ +theorem ContDiff.ftaylorSeries (hf : ContDiff 𝕜 n f) : + HasFTaylorSeriesUpTo n f (ftaylorSeries 𝕜 f) := by + simp only [← contDiffOn_univ, ← hasFTaylorSeriesUpToOn_univ_iff, ← ftaylorSeriesWithin_univ] + at hf ⊢ + exact ContDiffOn.ftaylorSeriesWithin hf uniqueDiffOn_univ -/-- When a function is `C^n` in a set `s` of unique differentiability, it admits -`ftaylorSeriesWithin 𝕜 f s` as a Taylor series up to order `n` in `s`. -/ +/-- For `n : ℕ∞`, a function is `C^n` iff it admits `ftaylorSeries 𝕜 f` +as a Taylor series up to order `n`. -/ theorem contDiff_iff_ftaylorSeries {n : ℕ∞} : ContDiff 𝕜 n f ↔ HasFTaylorSeriesUpTo n f (ftaylorSeries 𝕜 f) := by constructor · rw [← contDiffOn_univ, ← hasFTaylorSeriesUpToOn_univ_iff, ← ftaylorSeriesWithin_univ] - exact fun h => ContDiffOn.ftaylorSeriesWithin h uniqueDiffOn_univ - · intro h; exact ⟨ftaylorSeries 𝕜 f, h⟩ + exact fun h ↦ ContDiffOn.ftaylorSeriesWithin h uniqueDiffOn_univ + · exact fun h ↦ ⟨ftaylorSeries 𝕜 f, h⟩ theorem contDiff_iff_continuous_differentiable {n : ℕ∞} : ContDiff 𝕜 n f ↔ diff --git a/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean b/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean index a66bfc2242b36..45f5ee0897fc4 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean @@ -3,6 +3,7 @@ Copyright (c) 2019 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ +import Mathlib.Analysis.Calculus.FDeriv.Add import Mathlib.Analysis.Calculus.FDeriv.Equiv import Mathlib.Analysis.Calculus.FormalMultilinearSeries @@ -111,7 +112,7 @@ scoped [ContDiff] notation3 "ω" => (⊤ : WithTop ℕ∞) /-- Smoothness exponent for infinitely differentiable functions. -/ scoped [ContDiff] notation3 "∞" => ((⊤ : ℕ∞) : WithTop ℕ∞) -open scoped ContDiff +open scoped ContDiff Pointwise universe u uE uF @@ -514,6 +515,13 @@ lemma iteratedFDerivWithin_two_apply (f : E → F) {z : E} (hs : UniqueDiffOn simp only [iteratedFDerivWithin_succ_apply_right hs hz] rfl +/-- On a set of unique differentiability, the second derivative is obtained by taking the +derivative of the derivative. -/ +lemma iteratedFDerivWithin_two_apply' (f : E → F) {z : E} (hs : UniqueDiffOn 𝕜 s) (hz : z ∈ s) + (v w : E) : + iteratedFDerivWithin 𝕜 2 f s z ![v, w] = fderivWithin 𝕜 (fderivWithin 𝕜 f s) s z v w := + iteratedFDerivWithin_two_apply f hs hz _ + theorem Filter.EventuallyEq.iteratedFDerivWithin' (h : f₁ =ᶠ[𝓝[s] x] f) (ht : t ⊆ s) (n : ℕ) : iteratedFDerivWithin 𝕜 n f₁ t =ᶠ[𝓝[s] x] iteratedFDerivWithin 𝕜 n f t := by induction n with @@ -624,6 +632,48 @@ theorem HasFTaylorSeriesUpToOn.eq_iteratedFDerivWithin_of_uniqueDiffOn alias HasFTaylorSeriesUpToOn.eq_ftaylor_series_of_uniqueDiffOn := HasFTaylorSeriesUpToOn.eq_iteratedFDerivWithin_of_uniqueDiffOn +/-- The iterated derivative commutes with shifting the function by a constant on the left. -/ +lemma iteratedFDerivWithin_comp_add_left' (n : ℕ) (a : E) : + iteratedFDerivWithin 𝕜 n (fun z ↦ f (a + z)) s = + fun x ↦ iteratedFDerivWithin 𝕜 n f (a +ᵥ s) (a + x) := by + induction n with + | zero => simp [iteratedFDerivWithin] + | succ n IH => + ext v + rw [iteratedFDerivWithin_succ_eq_comp_left, iteratedFDerivWithin_succ_eq_comp_left] + simp only [Nat.succ_eq_add_one, IH, comp_apply, continuousMultilinearCurryLeftEquiv_symm_apply] + congr 2 + rw [fderivWithin_comp_add_left] + +/-- The iterated derivative commutes with shifting the function by a constant on the left. -/ +lemma iteratedFDerivWithin_comp_add_left (n : ℕ) (a : E) (x : E) : + iteratedFDerivWithin 𝕜 n (fun z ↦ f (a + z)) s x = + iteratedFDerivWithin 𝕜 n f (a +ᵥ s) (a + x) := by + simp [iteratedFDerivWithin_comp_add_left'] + +/-- The iterated derivative commutes with shifting the function by a constant on the right. -/ +lemma iteratedFDerivWithin_comp_add_right' (n : ℕ) (a : E) : + iteratedFDerivWithin 𝕜 n (fun z ↦ f (z + a)) s = + fun x ↦ iteratedFDerivWithin 𝕜 n f (a +ᵥ s) (x + a) := by + simpa [add_comm a] using iteratedFDerivWithin_comp_add_left' n a + +/-- The iterated derivative commutes with shifting the function by a constant on the right. -/ +lemma iteratedFDerivWithin_comp_add_right (n : ℕ) (a : E) (x : E) : + iteratedFDerivWithin 𝕜 n (fun z ↦ f (z + a)) s x = + iteratedFDerivWithin 𝕜 n f (a +ᵥ s) (x + a) := by + simp [iteratedFDerivWithin_comp_add_right'] + +/-- The iterated derivative commutes with subtracting a constant. -/ +lemma iteratedFDerivWithin_comp_sub' (n : ℕ) (a : E) : + iteratedFDerivWithin 𝕜 n (fun z ↦ f (z - a)) s = + fun x ↦ iteratedFDerivWithin 𝕜 n f (-a +ᵥ s) (x - a) := by + simpa [sub_eq_add_neg] using iteratedFDerivWithin_comp_add_right' n (-a) + +/-- The iterated derivative commutes with subtracting a constant. -/ +lemma iteratedFDerivWithin_comp_sub (n : ℕ) (a : E) : + iteratedFDerivWithin 𝕜 n (fun z ↦ f (z - a)) s x = + iteratedFDerivWithin 𝕜 n f (-a +ᵥ s) (x - a) := by + simp [iteratedFDerivWithin_comp_sub'] /-! ### Functions with a Taylor series on the whole space -/ @@ -867,3 +917,33 @@ lemma iteratedFDeriv_two_apply (f : E → F) (z : E) (m : Fin 2 → E) : iteratedFDeriv 𝕜 2 f z m = fderiv 𝕜 (fderiv 𝕜 f) z (m 0) (m 1) := by simp only [iteratedFDeriv_succ_apply_right] rfl + +/-- The iterated derivative commutes with shifting the function by a constant on the left. -/ +lemma iteratedFDeriv_comp_add_left' (n : ℕ) (a : E) : + iteratedFDeriv 𝕜 n (fun z ↦ f (a + z)) = fun x ↦ iteratedFDeriv 𝕜 n f (a + x) := by + simpa [← iteratedFDerivWithin_univ] using iteratedFDerivWithin_comp_add_left' n a (s := univ) + +/-- The iterated derivative commutes with shifting the function by a constant on the left. -/ +lemma iteratedFDeriv_comp_add_left (n : ℕ) (a : E) (x : E) : + iteratedFDeriv 𝕜 n (fun z ↦ f (a + z)) x = iteratedFDeriv 𝕜 n f (a + x) := by + simp [iteratedFDeriv_comp_add_left'] + +/-- The iterated derivative commutes with shifting the function by a constant on the right. -/ +lemma iteratedFDeriv_comp_add_right' (n : ℕ) (a : E) : + iteratedFDeriv 𝕜 n (fun z ↦ f (z + a)) = fun x ↦ iteratedFDeriv 𝕜 n f (x + a) := by + simpa [add_comm a] using iteratedFDeriv_comp_add_left' n a + +/-- The iterated derivative commutes with shifting the function by a constant on the right. -/ +lemma iteratedFDeriv_comp_add_right (n : ℕ) (a : E) (x : E) : + iteratedFDeriv 𝕜 n (fun z ↦ f (z + a)) x = iteratedFDeriv 𝕜 n f (x + a) := by + simp [iteratedFDeriv_comp_add_right'] + +/-- The iterated derivative commutes with subtracting a constant. -/ +lemma iteratedFDeriv_comp_sub' (n : ℕ) (a : E) : + iteratedFDeriv 𝕜 n (fun z ↦ f (z - a)) = fun x ↦ iteratedFDeriv 𝕜 n f (x - a) := by + simpa [sub_eq_add_neg] using iteratedFDeriv_comp_add_right' n (-a) + +/-- The iterated derivative commutes with subtracting a constant. -/ +lemma iteratedFDeriv_comp_sub (n : ℕ) (a : E) (x : E) : + iteratedFDeriv 𝕜 n (fun z ↦ f (z - a)) x = iteratedFDeriv 𝕜 n f (x - a) := by + simp [iteratedFDeriv_comp_sub'] diff --git a/Mathlib/Analysis/Calculus/FDeriv/Add.lean b/Mathlib/Analysis/Calculus/FDeriv/Add.lean index e513cd4e78b50..e6fb07b97cd41 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Add.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Add.lean @@ -711,4 +711,122 @@ theorem fderiv_const_sub (c : F) : fderiv 𝕜 (fun y => c - f y) x = -fderiv end Sub +section CompAdd + +/-! ### Derivative of the composition with a translation -/ + +open scoped Pointwise Topology + +theorem hasFDerivWithinAt_comp_add_right (a : E) : + HasFDerivWithinAt (fun x ↦ f (x + a)) f' s x ↔ HasFDerivWithinAt f f' (a +ᵥ s) (x + a) := by + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · have A : f = (fun x ↦ f (x + a)) ∘ (fun x ↦ x - a) := by ext; simp + rw [show x = (x + a) - a by abel] at h + rw [A] + have : HasFDerivWithinAt (fun x ↦ x - a) (ContinuousLinearMap.id 𝕜 E) (a +ᵥ s) (x + a) := by + simpa using (hasFDerivWithinAt_id (x + a) _).sub (hasFDerivWithinAt_const _ _ _) + apply h.comp (x + a) this (fun y hy ↦ ?_) + simpa [Set.mem_vadd_set_iff_neg_vadd_mem, add_comm, ← sub_eq_add_neg] using hy + · have : HasFDerivWithinAt (fun x ↦ x + a) (ContinuousLinearMap.id 𝕜 E) s x := by + simpa using (hasFDerivWithinAt_id x s (𝕜 := 𝕜)).add (hasFDerivWithinAt_const a x s (𝕜 := 𝕜)) + apply h.comp x this (fun y hy ↦ ?_) + simp [Set.mem_vadd_set_iff_neg_vadd_mem, hy] + +theorem differentiableWithinAt_comp_add_right (a : E) : + DifferentiableWithinAt 𝕜 (fun x ↦ f (x + a)) s x ↔ + DifferentiableWithinAt 𝕜 f (a +ᵥ s) (x + a) := by + simp [DifferentiableWithinAt, hasFDerivWithinAt_comp_add_right] + +theorem fderivWithin_comp_add_right (a : E) : + fderivWithin 𝕜 (fun x ↦ f (x + a)) s x = fderivWithin 𝕜 f (a +ᵥ s) (x + a) := by + simp only [fderivWithin, hasFDerivWithinAt_comp_add_right] + by_cases h : 𝓝[s \ {x}] x = ⊥ + · have h' : 𝓝[(a +ᵥ s) \ {x + a}] (x + a) = ⊥ := by + let e := Homeomorph.addRight a + have : 𝓝[(a +ᵥ s) \ {x + a}] (x + a) = map e (𝓝[s \ {x}] x) := by + rw [e.isEmbedding.map_nhdsWithin_eq] + congr + ext y + rw [← e.preimage_symm] + simp [e, Homeomorph.addRight, Set.mem_vadd_set_iff_neg_vadd_mem, add_comm] + rw [this, h, Filter.map_bot] + simp [h, h'] + · have h' : 𝓝[(a +ᵥ s) \ {x + a}] (x + a) ≠ ⊥ := by + intro h'' + let e := Homeomorph.addRight (-a) + have : 𝓝[s \ {x}] x = map e (𝓝[(a +ᵥ s) \ {x + a}] (x + a)) := by + rw [e.isEmbedding.map_nhdsWithin_eq] + congr + · simp [e] + ext y + rw [← e.preimage_symm] + simp [e, Homeomorph.addRight, Set.mem_vadd_set_iff_neg_vadd_mem, add_comm] + apply h + rw [this, h'', Filter.map_bot] + simp [h, h'] + +theorem hasFDerivWithinAt_comp_add_left (a : E) : + HasFDerivWithinAt (fun x ↦ f (a + x)) f' s x ↔ HasFDerivWithinAt f f' (a +ᵥ s) (a + x) := by + simpa [add_comm a] using hasFDerivWithinAt_comp_add_right a + +theorem differentiableWithinAt_comp_add_left (a : E) : + DifferentiableWithinAt 𝕜 (fun x ↦ f (a + x)) s x ↔ + DifferentiableWithinAt 𝕜 f (a +ᵥ s) (a + x) := by + simp [DifferentiableWithinAt, hasFDerivWithinAt_comp_add_left] + +theorem fderivWithin_comp_add_left (a : E) : + fderivWithin 𝕜 (fun x ↦ f (a + x)) s x = fderivWithin 𝕜 f (a +ᵥ s) (a + x) := by + simpa [add_comm a] using fderivWithin_comp_add_right a + +theorem hasFDerivAt_comp_add_right (a : E) : + HasFDerivAt (fun x ↦ f (x + a)) f' x ↔ HasFDerivAt f f' (x + a) := by + simp [← hasFDerivWithinAt_univ, hasFDerivWithinAt_comp_add_right] + +theorem differentiableAt_comp_add_right (a : E) : + DifferentiableAt 𝕜 (fun x ↦ f (x + a)) x ↔ DifferentiableAt 𝕜 f (x + a) := by + simp [DifferentiableAt, hasFDerivAt_comp_add_right] + +theorem fderiv_comp_add_right (a : E) : + fderiv 𝕜 (fun x ↦ f (x + a)) x = fderiv 𝕜 f (x + a) := by + simp [← fderivWithin_univ, fderivWithin_comp_add_right] + +theorem hasFDerivAt_comp_add_left (a : E) : + HasFDerivAt (fun x ↦ f (a + x)) f' x ↔ HasFDerivAt f f' (a + x) := by + simpa [add_comm a] using hasFDerivAt_comp_add_right a + +theorem differentiableAt_comp_add_left (a : E) : + DifferentiableAt 𝕜 (fun x ↦ f (a + x)) x ↔ DifferentiableAt 𝕜 f (a + x) := by + simp [DifferentiableAt, hasFDerivAt_comp_add_left] + +theorem fderiv_comp_add_left (a : E) : + fderiv 𝕜 (fun x ↦ f (a + x)) x = fderiv 𝕜 f (a + x) := by + simpa [add_comm a] using fderiv_comp_add_right a + +theorem hasFDerivWithinAt_comp_sub (a : E) : + HasFDerivWithinAt (fun x ↦ f (x - a)) f' s x ↔ HasFDerivWithinAt f f' (-a +ᵥ s) (x - a) := by + simpa [sub_eq_add_neg] using hasFDerivWithinAt_comp_add_right (-a) + +theorem differentiableWithinAt_comp_sub (a : E) : + DifferentiableWithinAt 𝕜 (fun x ↦ f (x - a)) s x ↔ + DifferentiableWithinAt 𝕜 f (-a +ᵥ s) (x - a) := by + simp [DifferentiableWithinAt, hasFDerivWithinAt_comp_sub] + +theorem fderivWithin_comp_sub (a : E) : + fderivWithin 𝕜 (fun x ↦ f (x - a)) s x = fderivWithin 𝕜 f (-a +ᵥ s) (x - a) := by + simpa [sub_eq_add_neg] using fderivWithin_comp_add_right (-a) + +theorem hasFDerivAt_comp_sub (a : E) : + HasFDerivAt (fun x ↦ f (x - a)) f' x ↔ HasFDerivAt f f' (x - a) := by + simp [← hasFDerivWithinAt_univ, hasFDerivWithinAt_comp_sub] + +theorem differentiableAt_comp_sub (a : E) : + DifferentiableAt 𝕜 (fun x ↦ f (x - a)) x ↔ DifferentiableAt 𝕜 f (x - a) := by + simp [DifferentiableAt, hasFDerivAt_comp_sub] + +theorem fderiv_comp_sub (a : E) : + fderiv 𝕜 (fun x ↦ f (x - a)) x = fderiv 𝕜 f (x - a) := by + simp [← fderivWithin_univ, fderivWithin_comp_sub] + +end CompAdd + end diff --git a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean index 48cb70ce773fd..d7b62c6e50328 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean @@ -222,6 +222,17 @@ protected theorem HasFPowerSeriesWithinOnBall.fderivWithin [CompleteSpace F] · simpa only [edist_eq_coe_nnnorm_sub, EMetric.mem_ball] using hz.2 · simpa using hz.1 +/-- If a function has a power series within a set on a ball, then so does its derivative. For a +version without completeness, but assuming that the function is analytic on the set `s`, see +`HasFPowerSeriesWithinOnBall.fderivWithin_of_mem_of_analyticOn`. -/ +protected theorem HasFPowerSeriesWithinOnBall.fderivWithin_of_mem [CompleteSpace F] + (h : HasFPowerSeriesWithinOnBall f p s x r) (hu : UniqueDiffOn 𝕜 s) (hx : x ∈ s) : + HasFPowerSeriesWithinOnBall (fderivWithin 𝕜 f s) p.derivSeries s x r := by + have : insert x s = s := insert_eq_of_mem hx + rw [← this] at hu + convert h.fderivWithin hu + exact this.symm + /-- If a function is analytic on a set `s`, so is its Fréchet derivative. -/ protected theorem AnalyticAt.fderiv [CompleteSpace F] (h : AnalyticAt 𝕜 f x) : AnalyticAt 𝕜 (fderiv 𝕜 f) x := by @@ -255,7 +266,7 @@ protected theorem AnalyticOnNhd.iteratedFDeriv [CompleteSpace F] (h : AnalyticOn simp @[deprecated (since := "2024-09-26")] -alias AnalyticOn.iteratedFDeriv := AnalyticOnNhd.iteratedFDeriv +protected 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 @@ -318,13 +329,13 @@ theorem HasFPowerSeriesWithinOnBall.hasSum_derivSeries_of_hasFDerivWithinAt 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, ?_⟩ +/-- If a function has a power series within a set on a ball, then so does its derivative. Version +assuming that the function is analytic on `s`. For a version without this assumption but requiring +that `F` is complete, see `HasFPowerSeriesWithinOnBall.fderivWithin_of_mem`. -/ +protected theorem HasFPowerSeriesWithinOnBall.fderivWithin_of_mem_of_analyticOn + (hr : HasFPowerSeriesWithinOnBall f p s x r) + (h : AnalyticOn 𝕜 f s) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) : + HasFPowerSeriesWithinOnBall (fderivWithin 𝕜 f s) p.derivSeries s x r := by 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 ⊢ @@ -332,6 +343,14 @@ protected theorem AnalyticOn.fderivWithin (h : AnalyticOn 𝕜 f s) (hu : Unique exact h.differentiableOn _ hy · rwa [insert_eq_of_mem hx] +/-- 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, hr.fderivWithin_of_mem_of_analyticOn h hu 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) diff --git a/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean b/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean index 2dfe679519c45..d95f0f210244f 100644 --- a/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean +++ b/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean @@ -75,6 +75,12 @@ theorem zero_apply (n : ℕ) : (0 : FormalMultilinearSeries 𝕜 E F) n = 0 := r @[simp] theorem neg_apply (f : FormalMultilinearSeries 𝕜 E F) (n : ℕ) : (-f) n = - f n := rfl +@[simp] +theorem add_apply (p q : FormalMultilinearSeries 𝕜 E F) (n : ℕ) : (p + q) n = p n + q n := rfl + +@[simp] +theorem sub_apply (p q : FormalMultilinearSeries 𝕜 E F) (n : ℕ) : (p - q) n = p n - q n := rfl + @[ext] protected theorem ext {p q : FormalMultilinearSeries 𝕜 E F} (h : ∀ n, p n = q n) : p = q := funext h @@ -125,6 +131,10 @@ theorem congr (p : FormalMultilinearSeries 𝕜 E F) {m n : ℕ} {v : Fin m → congr with ⟨i, hi⟩ exact h2 i hi hi +lemma congr_zero (p : FormalMultilinearSeries 𝕜 E F) {k l : ℕ} (h : k = l) (h' : p k = 0) : + p l = 0 := by + subst h; exact h' + /-- Composing each term `pₙ` in a formal multilinear series with `(u, ..., u)` where `u` is a fixed continuous linear map, gives a new formal multilinear series `p.compContinuousLinearMap u`. -/ def compContinuousLinearMap (p : FormalMultilinearSeries 𝕜 F G) (u : E →L[𝕜] F) : diff --git a/Mathlib/Data/Fintype/Card.lean b/Mathlib/Data/Fintype/Card.lean index 59e57330eb779..22e95987df0fd 100644 --- a/Mathlib/Data/Fintype/Card.lean +++ b/Mathlib/Data/Fintype/Card.lean @@ -706,15 +706,29 @@ theorem set_fintype_card_eq_univ_iff [Fintype α] (s : Set α) [Fintype s] : namespace Function.Embedding /-- An embedding from a `Fintype` to itself can be promoted to an equivalence. -/ -noncomputable def equivOfFintypeSelfEmbedding [Finite α] (e : α ↪ α) : α ≃ α := +noncomputable def equivOfFiniteSelfEmbedding [Finite α] (e : α ↪ α) : α ≃ α := Equiv.ofBijective e e.2.bijective_of_finite +@[deprecated (since := "2024-12-05")] +alias equivOfFintypeSelfEmbedding := equivOfFiniteSelfEmbedding + @[simp] -theorem equiv_of_fintype_self_embedding_to_embedding [Finite α] (e : α ↪ α) : - e.equivOfFintypeSelfEmbedding.toEmbedding = e := by +theorem toEmbedding_equivOfFiniteSelfEmbedding [Finite α] (e : α ↪ α) : + e.equivOfFiniteSelfEmbedding.toEmbedding = e := by ext rfl +@[deprecated (since := "2024-12-05")] +alias equiv_of_fintype_self_embedding_to_embedding := toEmbedding_equivOfFiniteSelfEmbedding + +/-- On a finite type, equivalence between the self-embeddings and the bijections. -/ +@[simps] noncomputable def _root_.Equiv.embeddingEquivOfFinite (α : Type*) [Finite α] : + (α ↪ α) ≃ (α ≃ α) where + toFun e := e.equivOfFiniteSelfEmbedding + invFun e := e.toEmbedding + left_inv e := rfl + right_inv e := by ext; rfl + /-- If `‖β‖ < ‖α‖` there are no embeddings `α ↪ β`. This is a formulation of the pigeonhole principle. @@ -750,7 +764,7 @@ end Function.Embedding @[simp] theorem Finset.univ_map_embedding {α : Type*} [Fintype α] (e : α ↪ α) : univ.map e = univ := by - rw [← e.equiv_of_fintype_self_embedding_to_embedding, univ_map_equiv_to_embedding] + rw [← e.toEmbedding_equivOfFiniteSelfEmbedding, univ_map_equiv_to_embedding] namespace Fintype diff --git a/Mathlib/Topology/Homeomorph.lean b/Mathlib/Topology/Homeomorph.lean index f98a10c4d8b6b..b7e93c00e5acf 100644 --- a/Mathlib/Topology/Homeomorph.lean +++ b/Mathlib/Topology/Homeomorph.lean @@ -198,6 +198,9 @@ theorem image_preimage (h : X ≃ₜ Y) (s : Set Y) : h '' (h ⁻¹' s) = s := theorem preimage_image (h : X ≃ₜ Y) (s : Set X) : h ⁻¹' (h '' s) = s := h.toEquiv.preimage_image s +theorem image_eq_preimage (h : X ≃ₜ Y) (s : Set X) : h '' s = h.symm ⁻¹' s := + h.toEquiv.image_eq_preimage s + lemma image_compl (h : X ≃ₜ Y) (s : Set X) : h '' (sᶜ) = (h '' s)ᶜ := h.toEquiv.image_compl s From 5d2254fa2ab034e34b4ae1363d48ea7acf430a6b Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Tue, 10 Dec 2024 08:25:53 +0000 Subject: [PATCH 628/829] feat: PrimeSpectrum has DiscreteTopology iff toLocalization is bijective (#19714) Third PR in the series #18891 and #19496 Main results: + `PrimeSpectrum.discreteTopology_iff_toPiLocalization_surjective`: the prime spectrum of a comm. semiring has discrete Zariski topology iff the map into the product of localizations at primes is surjective (equivalently, bijective). + `mapPiEvalRingHom_bijective`: the localization of a product ring $\prod_i R_i$ w.r.t. a "cylindrical" submonoid (of the form $\pi_i^{-1}(S)$, where S is a submonoid of some $R_i$) is canonically isomorphic to the localization of $R_i$ w.r.t. S. + `exists_maximal_nmem_range_sigmaToPi_of_infinite`: an infinite product of nontrivial rings has a maximal ideal that is not cylindrical. As a consequence, the map to the product of localizations at maximal ideals cannot be surjective (`MaximalSpectrum.toPiLocalization_not_surjective_of_infinite`), since the value at the non-cylindrical max ideal is determined by the values at the cylindrical max ideals. As a further consequence, if a ring is isomorphic to the product of its localizations at maximal ideals, the number of maximal ideals must be finite (`MaximalSpectrum.finite_of_toPiLocalization_surjective`). We rename `RingHom.toLocalizationIsMaximal` to `MaximalSpectrum.toPiLocalization` and similarly for associated lemmas. These were introduced in the previous two PRs and are relatively new, so I hope I won't need to add deprecations. --- .../PrimeSpectrum/Basic.lean | 51 ++-- Mathlib/RingTheory/Localization/AtPrime.lean | 35 +-- Mathlib/RingTheory/Localization/Basic.lean | 32 +++ Mathlib/RingTheory/MaximalSpectrum.lean | 220 +++++++++++++++++- Mathlib/RingTheory/PrimeSpectrum.lean | 53 +++++ 5 files changed, 355 insertions(+), 36 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index 2a631d1f70aeb..96a4c71a98a96 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -10,6 +10,7 @@ import Mathlib.RingTheory.KrullDimension.Basic import Mathlib.RingTheory.LocalRing.ResidueField.Defs import Mathlib.RingTheory.LocalRing.RingHom.Basic import Mathlib.RingTheory.Localization.Away.Basic +import Mathlib.RingTheory.MaximalSpectrum import Mathlib.Tactic.StacksAttribute import Mathlib.Topology.KrullDimension import Mathlib.Topology.Sober @@ -237,6 +238,12 @@ theorem discreteTopology_iff_finite_isMaximal_and_sInf_le_nilradical : rwa [← hM.eq_of_le hI.1 hMI] exact ⟨fin.subset hpm, hpm⟩ +theorem discreteTopology_of_toLocalization_surjective + (surj : Function.Surjective (toPiLocalization R)) : + DiscreteTopology (PrimeSpectrum R) := + discreteTopology_iff_finite_and_isPrime_imp_isMaximal.mpr ⟨finite_of_toPiLocalization_surjective + surj, fun I prime ↦ isMaximal_of_toPiLocalization_surjective surj ⟨I, prime⟩⟩ + section Comap variable {S' : Type*} [CommSemiring S'] @@ -542,11 +549,14 @@ theorem isLocalization_away_iff_atPrime_of_basicOpen_eq_singleton [Algebra R S] exact not_not.mpr (q.span_singleton_le_iff_mem.mp le) IsLocalization.isLocalization_iff_of_isLocalization _ _ (Localization.Away f) -variable [DiscreteTopology (PrimeSpectrum R)] +end BasicOpen + +section DiscreteTopology -variable (R) in -lemma _root_.RingHom.toLocalizationIsMaximal_surjective_of_discreteTopology : - Function.Surjective (RingHom.toLocalizationIsMaximal R) := fun x ↦ by +variable (R) [DiscreteTopology (PrimeSpectrum R)] + +theorem toPiLocalization_surjective_of_discreteTopology : + Function.Surjective (toPiLocalization R) := fun x ↦ by have (p : PrimeSpectrum R) : ∃ f, (basicOpen f : Set _) = {p} := have ⟨_, ⟨f, rfl⟩, hpf, hfp⟩ := isTopologicalBasis_basic_opens.isOpen_iff.mp (isOpen_discrete {p}) p rfl @@ -562,9 +572,8 @@ lemma _root_.RingHom.toLocalizationIsMaximal_surjective_of_discreteTopology : fun p _ ↦ TopologicalSpace.Opens.mem_iSup.mpr ⟨p, (hf p).ge rfl⟩ replace hf a : (basicOpen a.1 : Set _) = {e.symm a} := by simp_rw [← hf, Equiv.apply_ofInjective_symm] - have := (discreteTopology_iff_finite_and_isPrime_imp_isMaximal.mp ‹_›).2 obtain ⟨r, eq, -⟩ := Localization.existsUnique_algebraMap_eq_of_span_eq_top _ span_eq - (fun a ↦ algE a (x ⟨_, this _ inferInstance⟩)) fun a b ↦ by + (fun a ↦ algE a (x _)) fun a b ↦ by obtain rfl | ne := eq_or_ne a b; · rfl have ⟨n, hn⟩ : IsNilpotent (a * b : R) := (basicOpen_eq_bot_iff _).mp <| by simp_rw [basicOpen_mul, SetLike.ext'_iff, TopologicalSpace.Opens.coe_inf, hf] @@ -573,19 +582,35 @@ lemma _root_.RingHom.toLocalizationIsMaximal_surjective_of_discreteTopology : (S := Localization.Away (a * b : R)) <| hn ▸ ⟨n, rfl⟩ apply Subsingleton.elim refine ⟨r, funext fun I ↦ ?_⟩ - have := eq (e ⟨I, I.2.isPrime⟩) + have := eq (e I) rwa [← AlgEquiv.symm_apply_eq, AlgEquiv.commutes, e.symm_apply_apply] at this +theorem maximalSpectrumToPiLocalization_surjective_of_discreteTopology : + Function.Surjective (MaximalSpectrum.toPiLocalization R) := by + rw [← piLocalizationToMaximal_comp_toPiLocalization] + exact (piLocalizationToMaximal_surjective R).comp + (toPiLocalization_surjective_of_discreteTopology R) + /-- If the prime spectrum of a commutative semiring R has discrete Zariski topology, then R is canonically isomorphic to the product of its localizations at the (finitely many) maximal ideals. -/ @[stacks 00JA "See also `PrimeSpectrum.discreteTopology_iff_finite_isMaximal_and_sInf_le_nilradical`."] -def _root_.RingHom.toLocalizationIsMaximalEquiv : R ≃+* - Π I : {I : Ideal R // I.IsMaximal}, haveI : I.1.IsMaximal := I.2; Localization.AtPrime I.1 := - .ofBijective _ ⟨RingHom.toLocalizationIsMaximal_injective R, - RingHom.toLocalizationIsMaximal_surjective_of_discreteTopology R⟩ - -end BasicOpen +def MaximalSpectrum.toPiLocalizationEquivtoLocalizationEquiv : + R ≃+* MaximalSpectrum.PiLocalization R := + .ofBijective _ ⟨MaximalSpectrum.toPiLocalization_injective R, + maximalSpectrumToPiLocalization_surjective_of_discreteTopology R⟩ + +theorem discreteTopology_iff_toPiLocalization_surjective {R} [CommSemiring R] : + DiscreteTopology (PrimeSpectrum R) ↔ Function.Surjective (toPiLocalization R) := + ⟨fun _ ↦ toPiLocalization_surjective_of_discreteTopology _, + discreteTopology_of_toLocalization_surjective⟩ + +theorem discreteTopology_iff_toPiLocalization_bijective {R} [CommSemiring R] : + DiscreteTopology (PrimeSpectrum R) ↔ Function.Bijective (toPiLocalization R) := + discreteTopology_iff_toPiLocalization_surjective.trans + (and_iff_right <| toPiLocalization_injective _).symm + +end DiscreteTopology section Order diff --git a/Mathlib/RingTheory/Localization/AtPrime.lean b/Mathlib/RingTheory/Localization/AtPrime.lean index 1337c943de11e..e8d94b851a13e 100644 --- a/Mathlib/RingTheory/Localization/AtPrime.lean +++ b/Mathlib/RingTheory/Localization/AtPrime.lean @@ -249,24 +249,27 @@ theorem localRingHom_comp {S : Type*} [CommSemiring S] (J : Ideal S) [hJ : J.IsP localRingHom_unique _ _ _ _ fun r => by simp only [Function.comp_apply, RingHom.coe_comp, localRingHom_to_map] -end Localization +namespace AtPrime + +variable {ι : Type*} {R : ι → Type*} [∀ i, CommSemiring (R i)] +variable {i : ι} (I : Ideal (R i)) [I.IsPrime] -namespace RingHom +/-- `Localization.localRingHom` specialized to a projection homomorphism from a product ring. -/ +noncomputable abbrev mapPiEvalRingHom : + Localization.AtPrime (I.comap <| Pi.evalRingHom R i) →+* Localization.AtPrime I := + localRingHom _ _ _ rfl -variable (R) +theorem mapPiEvalRingHom_bijective : Function.Bijective (mapPiEvalRingHom I) := + Localization.mapPiEvalRingHom_bijective _ -/-- The canonical ring homomorphism from a commutative semiring to the product of its -localizations at all maximal ideals. It is always injective. -/ -def toLocalizationIsMaximal : R →+* - Π I : {I : Ideal R // I.IsMaximal}, haveI : I.1.IsMaximal := I.2; Localization.AtPrime I.1 := - Pi.ringHom fun _ ↦ algebraMap R _ +theorem mapPiEvalRingHom_comp_algebraMap : + (mapPiEvalRingHom I).comp (algebraMap _ _) = (algebraMap _ _).comp (Pi.evalRingHom R i) := + IsLocalization.map_comp _ -theorem toLocalizationIsMaximal_injective : - Function.Injective (RingHom.toLocalizationIsMaximal R) := fun r r' eq ↦ by - rw [← one_mul r, ← one_mul r'] - by_contra ne - have ⟨I, mI, hI⟩ := (Module.eqIdeal R r r').exists_le_maximal ((Ideal.ne_top_iff_one _).mpr ne) - have ⟨s, hs⟩ := (IsLocalization.eq_iff_exists I.primeCompl _).mp (congr_fun eq ⟨I, mI⟩) - exact s.2 (hI hs) +theorem mapPiEvalRingHom_algebraMap_apply {r : Π i, R i} : + mapPiEvalRingHom I (algebraMap _ _ r) = algebraMap _ _ (r i) := + localRingHom_to_map .. -end RingHom +end AtPrime + +end Localization diff --git a/Mathlib/RingTheory/Localization/Basic.lean b/Mathlib/RingTheory/Localization/Basic.lean index d2ed2970d4c2b..6f9bf16ab95c9 100644 --- a/Mathlib/RingTheory/Localization/Basic.lean +++ b/Mathlib/RingTheory/Localization/Basic.lean @@ -71,6 +71,38 @@ assert_not_exists Ideal open Function +namespace Localization + +open IsLocalization + +variable {ι : Type*} {R : ι → Type*} [∀ i, CommSemiring (R i)] +variable {i : ι} (S : Submonoid (R i)) + +/-- `IsLocalization.map` applied to a projection homomorphism from a product ring. -/ +noncomputable abbrev mapPiEvalRingHom : + Localization (S.comap <| Pi.evalRingHom R i) →+* Localization S := + map (T := S) _ (Pi.evalRingHom R i) le_rfl + +open Function in +theorem mapPiEvalRingHom_bijective : Bijective (mapPiEvalRingHom S) := by + let T := S.comap (Pi.evalRingHom R i) + classical + refine ⟨fun x₁ x₂ eq ↦ ?_, fun x ↦ ?_⟩ + · obtain ⟨r₁, s₁, rfl⟩ := mk'_surjective T x₁ + obtain ⟨r₂, s₂, rfl⟩ := mk'_surjective T x₂ + simp_rw [map_mk'] at eq + rw [IsLocalization.eq] at eq ⊢ + obtain ⟨s, hs⟩ := eq + refine ⟨⟨update 0 i s, by apply update_same i s.1 0 ▸ s.2⟩, funext fun j ↦ ?_⟩ + obtain rfl | ne := eq_or_ne j i + · simpa using hs + · simp [update_noteq ne] + · obtain ⟨r, s, rfl⟩ := mk'_surjective S x + exact ⟨mk' (M := T) _ (update 0 i r) ⟨update 0 i s, by apply update_same i s.1 0 ▸ s.2⟩, + by simp [map_mk']⟩ + +end Localization + section CommSemiring variable {R : Type*} [CommSemiring R] {M N : Submonoid R} {S : Type*} [CommSemiring S] diff --git a/Mathlib/RingTheory/MaximalSpectrum.lean b/Mathlib/RingTheory/MaximalSpectrum.lean index 959eae0a6633e..9e400e2b5fb1c 100644 --- a/Mathlib/RingTheory/MaximalSpectrum.lean +++ b/Mathlib/RingTheory/MaximalSpectrum.lean @@ -28,9 +28,7 @@ noncomputable section open scoped Classical -universe u v - -variable (R : Type u) [CommRing R] +variable (R S P : Type*) [CommSemiring R] [CommSemiring S] [CommSemiring P] /-- The maximal spectrum of a commutative ring `R` is the type of all maximal ideals of `R`. -/ @[ext] @@ -57,8 +55,8 @@ theorem toPrimeSpectrum_injective : (@toPrimeSpectrum R _).Injective := fun ⟨_ open PrimeSpectrum Set -variable (R) -variable [IsDomain R] (K : Type v) [Field K] [Algebra R K] [IsFractionRing R K] +variable (R : Type*) +variable [CommRing R] [IsDomain R] (K : Type*) [Field K] [Algebra R K] [IsFractionRing R K] /-- An integral domain is equal to the intersection of its localizations at all its maximal ideals viewed as subalgebras of its field of fractions. -/ @@ -82,8 +80,8 @@ end MaximalSpectrum namespace PrimeSpectrum -variable (R) -variable [IsDomain R] (K : Type v) [Field K] [Algebra R K] [IsFractionRing R K] +variable (R : Type*) +variable [CommRing R] [IsDomain R] (K : Type*) [Field K] [Algebra R K] [IsFractionRing R K] /-- An integral domain is equal to the intersection of its localizations at all its prime ideals viewed as subalgebras of its field of fractions. -/ @@ -93,3 +91,211 @@ theorem iInf_localization_eq_bot : ⨅ v : PrimeSpectrum R, simpa only [Algebra.mem_iInf] using fun hx ⟨v, hv⟩ ↦ hx ⟨v, hv.isPrime⟩ end PrimeSpectrum + +namespace MaximalSpectrum + +variable (R) + +/-- The product of localizations at all maximal ideals of a commutative semiring. -/ +abbrev PiLocalization : Type _ := Π I : MaximalSpectrum R, Localization.AtPrime I.1 + +/-- The canonical ring homomorphism from a commutative semiring to the product of its +localizations at all maximal ideals. It is always injective. -/ +def toPiLocalization : R →+* PiLocalization R := algebraMap R _ + +theorem toPiLocalization_injective : Function.Injective (toPiLocalization R) := fun r r' eq ↦ by + rw [← one_mul r, ← one_mul r'] + by_contra ne + have ⟨I, mI, hI⟩ := (Module.eqIdeal R r r').exists_le_maximal ((Ideal.ne_top_iff_one _).mpr ne) + have ⟨s, hs⟩ := (IsLocalization.eq_iff_exists I.primeCompl _).mp (congr_fun eq ⟨I, mI⟩) + exact s.2 (hI hs) + +theorem toPiLocalization_apply_apply {r I} : toPiLocalization R r I = algebraMap R _ r := rfl + +variable {R S} (f : R →+* S) (g : S →+* P) (hf : Function.Bijective f) (hg : Function.Bijective g) + +/-- Functoriality of `PiLocalization` but restricted to bijective ring homs. +If R and S are commutative rings, surjectivity would be enough. -/ +noncomputable def mapPiLocalization : PiLocalization R →+* PiLocalization S := + Pi.ringHom fun I ↦ (Localization.localRingHom _ _ f rfl).comp <| + Pi.evalRingHom _ (⟨_, I.2.comap_bijective f hf⟩ : MaximalSpectrum R) + +theorem mapPiLocalization_naturality : + (mapPiLocalization f hf).comp (toPiLocalization R) = + (toPiLocalization S).comp f := by + ext r I + show Localization.localRingHom _ _ _ rfl (algebraMap _ _ r) = algebraMap _ _ (f r) + simp_rw [← IsLocalization.mk'_one (M := (I.1.comap f).primeCompl), Localization.localRingHom_mk', + ← IsLocalization.mk'_one (M := I.1.primeCompl), Submonoid.coe_one, map_one f] + rfl + +theorem mapPiLocalization_id : mapPiLocalization (.id R) Function.bijective_id = .id _ := + RingHom.ext fun _ ↦ funext fun _ ↦ congr($(Localization.localRingHom_id _) _) + +theorem mapPiLocalization_comp : + mapPiLocalization (g.comp f) (hg.comp hf) = + (mapPiLocalization g hg).comp (mapPiLocalization f hf) := + RingHom.ext fun _ ↦ funext fun _ ↦ congr($(Localization.localRingHom_comp _ _ _ _ rfl _ rfl) _) + +theorem mapPiLocalization_bijective : Function.Bijective (mapPiLocalization f hf) := by + let f := RingEquiv.ofBijective f hf + let e := RingEquiv.ofRingHom (mapPiLocalization f hf) + (mapPiLocalization (f.symm : S →+* R) f.symm.bijective) ?_ ?_ + · exact e.bijective + · rw [← mapPiLocalization_comp] + simp_rw [RingEquiv.comp_symm, mapPiLocalization_id] + · rw [← mapPiLocalization_comp] + simp_rw [RingEquiv.symm_comp, mapPiLocalization_id] + +section Pi + +variable {ι} (R : ι → Type*) [∀ i, CommSemiring (R i)] [∀ i, Nontrivial (R i)] + +theorem toPiLocalization_not_surjective_of_infinite [Infinite ι] : + ¬ Function.Surjective (toPiLocalization (Π i, R i)) := fun surj ↦ by + have ⟨J, max, nmem⟩ := PrimeSpectrum.exists_maximal_nmem_range_sigmaToPi_of_infinite R + obtain ⟨r, hr⟩ := surj (Function.update 0 ⟨J, max⟩ 1) + have : r = 0 := funext fun i ↦ toPiLocalization_injective _ <| funext fun I ↦ by + replace hr := congr_fun hr ⟨_, I.2.comap_piEvalRingHom⟩ + dsimp only [toPiLocalization_apply_apply, Subtype.coe_mk] at hr + simp_rw [toPiLocalization_apply_apply, + ← Localization.AtPrime.mapPiEvalRingHom_algebraMap_apply, hr] + rw [Function.update_noteq]; · simp_rw [Pi.zero_apply, map_zero] + exact fun h ↦ nmem ⟨⟨i, I.1, I.2.isPrime⟩, PrimeSpectrum.ext congr($h.1)⟩ + replace hr := congr_fun hr ⟨J, max⟩ + rw [this, map_zero, Function.update_same] at hr + exact zero_ne_one hr + +variable {R} + +theorem finite_of_toPiLocalization_pi_surjective + (h : Function.Surjective (toPiLocalization (Π i, R i))) : + Finite ι := by + contrapose h; rw [not_finite_iff_infinite] at h + exact toPiLocalization_not_surjective_of_infinite _ + +end Pi + +theorem finite_of_toPiLocalization_surjective + (surj : Function.Surjective (toPiLocalization R)) : + Finite (MaximalSpectrum R) := by + replace surj := mapPiLocalization_bijective _ ⟨toPiLocalization_injective R, surj⟩ + |>.2.comp surj + rw [← RingHom.coe_comp, mapPiLocalization_naturality, RingHom.coe_comp] at surj + exact finite_of_toPiLocalization_pi_surjective surj.of_comp + +end MaximalSpectrum + +namespace PrimeSpectrum + +variable (R) + +/-- The product of localizations at all prime ideals of a commutative semiring. -/ +abbrev PiLocalization : Type _ := Π p : PrimeSpectrum R, Localization p.asIdeal.primeCompl + +/-- The canonical ring homomorphism from a commutative semiring to the product of its +localizations at all prime ideals. It is always injective. -/ +def toPiLocalization : R →+* PiLocalization R := algebraMap R _ + +theorem toPiLocalization_injective : Function.Injective (toPiLocalization R) := + fun _ _ eq ↦ MaximalSpectrum.toPiLocalization_injective R <| + funext fun I ↦ congr_fun eq I.toPrimeSpectrum + +/-- The projection from the product of localizations at primes to the product of +localizations at maximal ideals. -/ +def piLocalizationToMaximal : PiLocalization R →+* MaximalSpectrum.PiLocalization R := + Pi.ringHom fun I ↦ Pi.evalRingHom _ I.toPrimeSpectrum + +theorem piLocalizationToMaximal_surjective : Function.Surjective (piLocalizationToMaximal R) := + fun r ↦ ⟨fun I ↦ if h : I.1.IsMaximal then r ⟨_, h⟩ else 0, funext fun _ ↦ dif_pos _⟩ + +variable {R} + +/-- If R has Krull dimension ≤ 0, then `piLocalizationToIsMaximal R` is an isomorphism. -/ +def piLocalizationToMaximalEquiv (h : ∀ I : Ideal R, I.IsPrime → I.IsMaximal) : + PiLocalization R ≃+* MaximalSpectrum.PiLocalization R where + __ := piLocalizationToMaximal R + invFun := Pi.ringHom fun I ↦ Pi.evalRingHom _ (⟨_, h _ I.2⟩ : MaximalSpectrum R) + left_inv _ := rfl + right_inv _ := rfl + +theorem piLocalizationToMaximal_bijective (h : ∀ I : Ideal R, I.IsPrime → I.IsMaximal) : + Function.Bijective (piLocalizationToMaximal R) := + (piLocalizationToMaximalEquiv h).bijective + +theorem piLocalizationToMaximal_comp_toPiLocalization : + (piLocalizationToMaximal R).comp (toPiLocalization R) = MaximalSpectrum.toPiLocalization R := + rfl + +variable {S} + +theorem isMaximal_of_toPiLocalization_surjective (surj : Function.Surjective (toPiLocalization R)) + (I : PrimeSpectrum R) : I.1.IsMaximal := by + have ⟨J, max, le⟩ := I.1.exists_le_maximal I.2.ne_top + obtain ⟨r, hr⟩ := surj (Function.update 0 ⟨J, max.isPrime⟩ 1) + by_contra h + have hJ : algebraMap _ _ r = _ := (congr_fun hr _).trans (Function.update_same _ _ _) + have hI : algebraMap _ _ r = _ := congr_fun hr I + rw [← IsLocalization.lift_eq (M := J.primeCompl) (S := Localization J.primeCompl), hJ, map_one, + Function.update_noteq] at hI + · exact one_ne_zero hI + · intro eq; have : I.1 = J := congr_arg (·.1) eq; exact h (this ▸ max) + · exact fun ⟨s, hs⟩ ↦ IsLocalization.map_units (M := I.1.primeCompl) _ ⟨s, fun h ↦ hs (le h)⟩ + +variable (f : R →+* S) + +/-- A ring homomorphism induces a homomorphism between the products of localizations at primes. -/ +noncomputable def mapPiLocalization : PiLocalization R →+* PiLocalization S := + Pi.ringHom fun I ↦ (Localization.localRingHom _ I.1 f rfl).comp (Pi.evalRingHom _ (f.specComap I)) + +theorem mapPiLocalization_naturality : + (mapPiLocalization f).comp (toPiLocalization R) = (toPiLocalization S).comp f := by + ext r I + show Localization.localRingHom _ _ _ rfl (algebraMap _ _ r) = algebraMap _ _ (f r) + simp_rw [← IsLocalization.mk'_one (M := (I.1.comap f).primeCompl), Localization.localRingHom_mk', + ← IsLocalization.mk'_one (M := I.1.primeCompl), Submonoid.coe_one, map_one f] + rfl + +theorem mapPiLocalization_id : mapPiLocalization (.id R) = .id _ := by + ext; exact congr($(Localization.localRingHom_id _) _) + +theorem mapPiLocalization_comp (g : S →+* P) : + mapPiLocalization (g.comp f) = (mapPiLocalization g).comp (mapPiLocalization f) := by + ext; exact congr($(Localization.localRingHom_comp _ _ _ _ rfl _ rfl) _) + +theorem mapPiLocalization_bijective (hf : Function.Bijective f) : + Function.Bijective (mapPiLocalization f) := by + let f := RingEquiv.ofBijective f hf + let e := RingEquiv.ofRingHom (mapPiLocalization (f : R →+* S)) (mapPiLocalization f.symm) ?_ ?_ + · exact e.bijective + · rw [← mapPiLocalization_comp, RingEquiv.comp_symm, mapPiLocalization_id] + · rw [← mapPiLocalization_comp, RingEquiv.symm_comp, mapPiLocalization_id] + +section Pi + +variable {ι} (R : ι → Type*) [∀ i, CommSemiring (R i)] [∀ i, Nontrivial (R i)] + +theorem toPiLocalization_not_surjective_of_infinite [Infinite ι] : + ¬ Function.Surjective (toPiLocalization (Π i, R i)) := + fun surj ↦ MaximalSpectrum.toPiLocalization_not_surjective_of_infinite R <| by + rw [← piLocalizationToMaximal_comp_toPiLocalization] + exact (piLocalizationToMaximal_surjective _).comp surj + +variable {R} + +theorem finite_of_toPiLocalization_pi_surjective + (h : Function.Surjective (toPiLocalization (Π i, R i))) : + Finite ι := by + contrapose h; rw [not_finite_iff_infinite] at h + exact toPiLocalization_not_surjective_of_infinite _ + +end Pi + +theorem finite_of_toPiLocalization_surjective + (surj : Function.Surjective (toPiLocalization R)) : + Finite (PrimeSpectrum R) := by + replace surj := (mapPiLocalization_bijective _ ⟨toPiLocalization_injective R, surj⟩).2.comp surj + rw [← RingHom.coe_comp, mapPiLocalization_naturality, RingHom.coe_comp] at surj + exact finite_of_toPiLocalization_pi_surjective surj.of_comp + +end PrimeSpectrum diff --git a/Mathlib/RingTheory/PrimeSpectrum.lean b/Mathlib/RingTheory/PrimeSpectrum.lean index 5ac44c10a7f5f..f4f648dbb34cc 100644 --- a/Mathlib/RingTheory/PrimeSpectrum.lean +++ b/Mathlib/RingTheory/PrimeSpectrum.lean @@ -563,6 +563,59 @@ theorem localization_specComap_range [Algebra R S] (M : Submonoid R) [IsLocaliza ext1 exact IsLocalization.comap_map_of_isPrime_disjoint M S _ x.2 h +section Pi + +variable {ι} (R : ι → Type*) [∀ i, CommSemiring (R i)] + +/-- The canonical map from a disjoint union of prime spectra of commutative semirings to +the prime spectrum of the product semiring. -/ +/- TODO: show this is always a topological embedding (even when ι is infinite) +and is a homeomorphism when ι is finite. -/ +@[simps] def sigmaToPi : (Σ i, PrimeSpectrum (R i)) → PrimeSpectrum (Π i, R i) + | ⟨i, p⟩ => (Pi.evalRingHom R i).specComap p + +theorem sigmaToPi_injective : (sigmaToPi R).Injective := fun ⟨i, p⟩ ⟨j, q⟩ eq ↦ by + obtain rfl | ne := eq_or_ne i j + · congr; ext x + simpa using congr_arg (Function.update (0 : ∀ i, R i) i x ∈ ·.asIdeal) eq + · refine (p.1.ne_top_iff_one.mp p.2.ne_top ?_).elim + have : Function.update (1 : ∀ i, R i) j 0 ∈ (sigmaToPi R ⟨j, q⟩).asIdeal := by simp + simpa [← eq, Function.update_noteq ne] + +variable [Infinite ι] [∀ i, Nontrivial (R i)] + +/-- An infinite product of nontrivial commutative semirings has a maximal ideal outside of the +range of `sigmaToPi`, i.e. is not of the form `πᵢ⁻¹(𝔭)` for some prime `𝔭 ⊂ R i`, where +`πᵢ : (Π i, R i) →+* R i` is the projection. For a complete description of all prime ideals, +see https://math.stackexchange.com/a/1563190. -/ +theorem exists_maximal_nmem_range_sigmaToPi_of_infinite : + ∃ (I : Ideal (Π i, R i)) (_ : I.IsMaximal), ⟨I, inferInstance⟩ ∉ Set.range (sigmaToPi R) := by + let J : Ideal (Π i, R i) := -- `J := Π₀ i, R i` is an ideal in `Π i, R i` + { __ := AddMonoidHom.mrange DFinsupp.coeFnAddMonoidHom + smul_mem' := by + rintro r _ ⟨x, rfl⟩ + refine ⟨.mk x.support fun i ↦ r i * x i, funext fun i ↦ show dite _ _ _ = _ from ?_⟩ + simp_rw [DFinsupp.coeFnAddMonoidHom] + refine dite_eq_left_iff.mpr fun h ↦ ?_ + rw [DFinsupp.not_mem_support_iff.mp h, mul_zero] } + have ⟨I, max, le⟩ := J.exists_le_maximal <| (Ideal.ne_top_iff_one _).mpr <| by + -- take a maximal ideal I containing J + rintro ⟨x, hx⟩ + have ⟨i, hi⟩ := x.support.exists_not_mem + simpa [DFinsupp.coeFnAddMonoidHom, DFinsupp.not_mem_support_iff.mp hi] using congr_fun hx i + refine ⟨I, max, fun ⟨⟨i, p⟩, eq⟩ ↦ ?_⟩ + -- then I is not in the range of `sigmaToPi` + have : ⇑(DFinsupp.single i 1) ∉ (sigmaToPi R ⟨i, p⟩).asIdeal := by + simpa using p.1.ne_top_iff_one.mp p.2.ne_top + rw [eq] at this + exact this (le ⟨.single i 1, rfl⟩) + +theorem sigmaToPi_not_surjective_of_infinite : ¬ (sigmaToPi R).Surjective := fun surj ↦ + have ⟨_, _, nmem⟩ := exists_maximal_nmem_range_sigmaToPi_of_infinite R + (Set.range_eq_univ.mpr surj ▸ nmem) ⟨⟩ + +end Pi + end PrimeSpectrum section SpecOfSurjective From 229a92cd1b62354d83b60d88344bb54bc7c214c6 Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 10 Dec 2024 10:16:51 +0000 Subject: [PATCH 629/829] feat: `#trans_imports` shows the imports of the current module (#19129) `#trans_imports` reports how many transitive imports the current module has. The command takes two optional inputs: * `#trans_imports str`, where `str` is a `String` also shows the transitively imported modules whose name begins with `str`; * `#trans_imports (str)? at most n` behaves like `#trans_imports (str)?`, except that instead of reporting the *exact* count of transitive imports, it either confirms that their number is at most `n`, or warns that it is larger. The `at_most n` option is useful for tests, but could potentially be used to limit import growth in mathlib. [Zulip discussion](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/import.20changes.20bot.20on.20PRs/near/482768569) --- Mathlib.lean | 1 + Mathlib/Tactic/Common.lean | 1 + Mathlib/Util/TransImports.lean | 46 ++++++++++++++++++++++++++++++++++ MathlibTest/TransImports.lean | 10 ++++++++ 4 files changed, 58 insertions(+) create mode 100644 Mathlib/Util/TransImports.lean create mode 100644 MathlibTest/TransImports.lean diff --git a/Mathlib.lean b/Mathlib.lean index 59289e49c814b..9d04fd689ce22 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -5415,5 +5415,6 @@ import Mathlib.Util.Superscript import Mathlib.Util.SynthesizeUsing import Mathlib.Util.Tactic import Mathlib.Util.TermBeta +import Mathlib.Util.TransImports import Mathlib.Util.WhatsNew import Mathlib.Util.WithWeakNamespace diff --git a/Mathlib/Tactic/Common.lean b/Mathlib/Tactic/Common.lean index 3c61c4a884912..94a9e57234ef3 100644 --- a/Mathlib/Tactic/Common.lean +++ b/Mathlib/Tactic/Common.lean @@ -112,6 +112,7 @@ import Mathlib.Tactic.Widget.Conv import Mathlib.Tactic.WLOG import Mathlib.Util.AssertExists import Mathlib.Util.CountHeartbeats +import Mathlib.Util.TransImports import Mathlib.Util.WhatsNew /-! diff --git a/Mathlib/Util/TransImports.lean b/Mathlib/Util/TransImports.lean new file mode 100644 index 0000000000000..701b77fbf8bc3 --- /dev/null +++ b/Mathlib/Util/TransImports.lean @@ -0,0 +1,46 @@ +/- +Copyright (c) 2024 Damiano Testa. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Damiano Testa +-/ + +import Mathlib.Init + +/-! # The `#trans_imports` command + +`#trans_imports` reports how many transitive imports the current module has. +The command takes an optional string input: `#trans_imports str` also shows the transitively +imported modules whose name begins with `str`. +-/ + +/-- +`#trans_imports` reports how many transitive imports the current module has. +The command takes an optional string input: `#trans_imports str` also shows the transitively +imported modules whose name begins with `str`. + +Mostly for the sake of tests, the command also takes an optional `at_most x` input: +if the number of imports does not exceed `x`, then the message involves `x`, rather than the +actual, possibly varying, number of imports. +-/ +syntax (name := transImportsStx) "#trans_imports" (ppSpace str)? (&" at_most " num)? : command + +open Lean in +@[inherit_doc transImportsStx] +elab_rules : command +| `(command| #trans_imports $(stx)? $[at_most $le:num]?) => do + let imports := (← getEnv).allImportedModuleNames + let currMod := if let mod@(.str _ _) := ← getMainModule then m!"'{mod}' has " else "" + let rest := match stx with + | none => m!"" + | some str => + let imps := imports.filterMap fun (i : Name) => + if i.toString.startsWith str.getString then some i else none + m!"\n\n{imps.size} starting with {str}:\n{imps.qsort (·.toString < ·.toString)}" + match le with + | none => logInfo m!"{currMod}{imports.size} transitive imports{rest}" + | some bd => + if imports.size ≤ bd.getNat then + logInfo m!"{currMod}at most {bd} transitive imports{rest}" + else + logWarningAt bd + m!"{currMod}{imports.size} transitive imports, exceeding the expected bound of {bd}{rest}" diff --git a/MathlibTest/TransImports.lean b/MathlibTest/TransImports.lean new file mode 100644 index 0000000000000..964187009033d --- /dev/null +++ b/MathlibTest/TransImports.lean @@ -0,0 +1,10 @@ +import Mathlib.Util.TransImports + +/-- +info: 'MathlibTest.TransImports' has at most 500 transitive imports + +3 starting with "Lean.Elab.I": +[Lean.Elab.InfoTree, Lean.Elab.InfoTree.Main, Lean.Elab.InfoTree.Types] +-/ +#guard_msgs in +#trans_imports "Lean.Elab.I" at_most 500 From 6ae27fd154da91d50187a032da5fb41c3fd05577 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 10 Dec 2024 11:05:02 +0000 Subject: [PATCH 630/829] chore(Order/Ideal): define `Order.Cofinal` in terms of `IsCofinal` (#19681) --- Mathlib/ModelTheory/PartialEquiv.lean | 4 ++-- Mathlib/Order/CountableDenseLinearOrder.lean | 6 +++--- Mathlib/Order/Ideal.lean | 15 +++++++++------ 3 files changed, 14 insertions(+), 11 deletions(-) diff --git a/Mathlib/ModelTheory/PartialEquiv.lean b/Mathlib/ModelTheory/PartialEquiv.lean index 52f75fe5211e9..6a0442f49b37a 100644 --- a/Mathlib/ModelTheory/PartialEquiv.lean +++ b/Mathlib/ModelTheory/PartialEquiv.lean @@ -451,13 +451,13 @@ protected alias ⟨cod, _⟩ := isExtensionPair_iff_cod def definedAtLeft (h : L.IsExtensionPair M N) (m : M) : Order.Cofinal (FGEquiv L M N) where carrier := {f | m ∈ f.val.dom} - mem_gt := fun f => h f m + isCofinal := fun f => h f m /-- The cofinal set of finite equivalences with a given element in their codomain. -/ def definedAtRight (h : L.IsExtensionPair N M) (n : N) : Order.Cofinal (FGEquiv L M N) where carrier := {f | n ∈ f.val.cod} - mem_gt := fun f => h.cod f n + isCofinal := fun f => h.cod f n end IsExtensionPair diff --git a/Mathlib/Order/CountableDenseLinearOrder.lean b/Mathlib/Order/CountableDenseLinearOrder.lean index 52b1a34e4be76..31a60023ff36c 100644 --- a/Mathlib/Order/CountableDenseLinearOrder.lean +++ b/Mathlib/Order/CountableDenseLinearOrder.lean @@ -170,7 +170,7 @@ variable (β) def definedAtLeft [DenselyOrdered β] [NoMinOrder β] [NoMaxOrder β] [Nonempty β] (a : α) : Cofinal (PartialIso α β) where carrier := {f | ∃ b : β, (a, b) ∈ f.val} - mem_gt f := by + isCofinal f := by cases' exists_across f a with b a_b refine ⟨⟨insert (a, b) f.val, fun p hp q hq ↦ ?_⟩, ⟨b, Finset.mem_insert_self _ _⟩, @@ -190,8 +190,8 @@ variable (α) {β} def definedAtRight [DenselyOrdered α] [NoMinOrder α] [NoMaxOrder α] [Nonempty α] (b : β) : Cofinal (PartialIso α β) where carrier := {f | ∃ a, (a, b) ∈ f.val} - mem_gt f := by - rcases (definedAtLeft α b).mem_gt f.comm with ⟨f', ⟨a, ha⟩, hl⟩ + isCofinal f := by + rcases (definedAtLeft α b).isCofinal f.comm with ⟨f', ⟨a, ha⟩, hl⟩ refine ⟨f'.comm, ⟨a, ?_⟩, ?_⟩ · change (a, b) ∈ f'.val.image _ rwa [← Finset.mem_coe, Finset.coe_image, Equiv.image_eq_preimage] diff --git a/Mathlib/Order/Ideal.lean b/Mathlib/Order/Ideal.lean index 30f3f80385924..7603c89a04604 100644 --- a/Mathlib/Order/Ideal.lean +++ b/Mathlib/Order/Ideal.lean @@ -6,6 +6,7 @@ Authors: David Wärn import Mathlib.Logic.Encodable.Basic import Mathlib.Order.Atoms import Mathlib.Order.Chain +import Mathlib.Order.Cofinal import Mathlib.Order.UpperLower.Basic import Mathlib.Data.Set.Subsingleton @@ -463,15 +464,17 @@ structure Cofinal (P) [Preorder P] where /-- The carrier of a `Cofinal` is the underlying set. -/ carrier : Set P /-- The `Cofinal` contains arbitrarily large elements. -/ - mem_gt : ∀ x : P, ∃ y ∈ carrier, x ≤ y + isCofinal : IsCofinal carrier + +@[deprecated Cofinal.isCofinal (since := "2024-12-02")] +alias Cofinal.mem_gt := Cofinal.isCofinal namespace Cofinal variable [Preorder P] instance : Inhabited (Cofinal P) := - ⟨{ carrier := univ - mem_gt := fun x ↦ ⟨x, trivial, le_rfl⟩ }⟩ + ⟨_, .univ⟩ instance : Membership P (Cofinal P) := ⟨fun D x ↦ x ∈ D.carrier⟩ @@ -480,13 +483,13 @@ variable (D : Cofinal P) (x : P) /-- A (noncomputable) element of a cofinal set lying above a given element. -/ noncomputable def above : P := - Classical.choose <| D.mem_gt x + Classical.choose <| D.isCofinal x theorem above_mem : D.above x ∈ D := - (Classical.choose_spec <| D.mem_gt x).1 + (Classical.choose_spec <| D.isCofinal x).1 theorem le_above : x ≤ D.above x := - (Classical.choose_spec <| D.mem_gt x).2 + (Classical.choose_spec <| D.isCofinal x).2 end Cofinal From 99672f302276295ff76577128f36894038f35394 Mon Sep 17 00:00:00 2001 From: Vasily Nesterov Date: Tue, 10 Dec 2024 11:05:05 +0000 Subject: [PATCH 631/829] feat(Analysis/Analytic): uniqueness of function series converges to (#19851) Prove `HasFPowerSeriesOnBall.unique` (and its `HasFPowerSeriesWithinOnBall` variation): if two functions have the same series expansion, they coincide on the ball of convergence. It is direct translation from `HasSum.unique`. --- Mathlib/Analysis/Analytic/Basic.lean | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/Mathlib/Analysis/Analytic/Basic.lean b/Mathlib/Analysis/Analytic/Basic.lean index 3ebd74ae9c5f3..e1e5359b0a1de 100644 --- a/Mathlib/Analysis/Analytic/Basic.lean +++ b/Mathlib/Analysis/Analytic/Basic.lean @@ -530,6 +530,15 @@ theorem HasFPowerSeriesAt.congr (hf : HasFPowerSeriesAt f p x) (hg : f =ᶠ[𝓝 (h₁.mono (lt_min h₁.r_pos h₂pos) inf_le_left).congr fun y hy => h₂ (EMetric.ball_subset_ball inf_le_right hy)⟩ +theorem HasFPowerSeriesWithinOnBall.unique (hf : HasFPowerSeriesWithinOnBall f p s x r) + (hg : HasFPowerSeriesWithinOnBall g p s x r) : + (insert x s ∩ EMetric.ball x r).EqOn f g := fun _ hy ↦ + (hf.hasSum_sub hy).unique (hg.hasSum_sub hy) + +theorem HasFPowerSeriesOnBall.unique (hf : HasFPowerSeriesOnBall f p x r) + (hg : HasFPowerSeriesOnBall g p x r) : (EMetric.ball x r).EqOn f g := fun _ hy ↦ + (hf.hasSum_sub hy).unique (hg.hasSum_sub hy) + protected theorem HasFPowerSeriesWithinAt.eventually (hf : HasFPowerSeriesWithinAt f p s x) : ∀ᶠ r : ℝ≥0∞ in 𝓝[>] 0, HasFPowerSeriesWithinOnBall f p s x r := let ⟨_, hr⟩ := hf From bfdd6039bf11a194e78a351b27e21774684a016e Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 10 Dec 2024 11:05:06 +0000 Subject: [PATCH 632/829] chore: backports from leanprover/lean4#6123 (#19852) leanprover/lean4#6123 fixes some inconsistencies in the behaviour of `simp`, by propagating `Simp.Config` settings through to `Meta.Config`. This requires us to be more explicit about unfolding `let` bindings in many places in Mathlib. Co-authored-by: Johan Commelin --- .../Category/ModuleCat/ChangeOfRings.lean | 2 +- .../ModuleCat/Differentials/Basic.lean | 4 ++-- .../Algebra/Homology/BifunctorAssociator.lean | 7 ++++--- .../Homology/DerivedCategory/Ext/Basic.lean | 14 ++++++------- .../Homology/ShortComplex/Abelian.lean | 5 +++-- Mathlib/Algebra/Lie/CartanExists.lean | 2 +- Mathlib/Algebra/Lie/Weights/Chain.lean | 8 ++++---- Mathlib/Algebra/Lie/Weights/Killing.lean | 4 ++-- .../Algebra/Module/FinitePresentation.lean | 2 +- Mathlib/Algebra/Module/FreeLocus.lean | 8 ++++---- .../Algebra/Module/LinearMap/Polynomial.lean | 4 ++-- Mathlib/Algebra/Order/Sub/WithTop.lean | 3 ++- Mathlib/Algebra/Polynomial/Div.lean | 2 +- .../EllipticCurve/Group.lean | 6 +++--- .../Morphisms/SurjectiveOnStalks.lean | 9 ++++++--- .../Morphisms/UniversallyClosed.lean | 2 +- .../PrimeSpectrum/Basic.lean | 4 ++-- Mathlib/AlgebraicGeometry/RationalMap.lean | 4 ++-- .../Sites/MorphismProperty.lean | 2 +- .../AlgebraicGeometry/ValuativeCriterion.lean | 4 ++-- .../AlternatingFaceMapComplex.lean | 8 ++++---- .../DoldKan/Decomposition.lean | 2 +- Mathlib/Analysis/Analytic/Inverse.lean | 6 +++--- .../Analysis/BoxIntegral/UnitPartition.lean | 5 +++-- .../Instances.lean | 2 +- .../BumpFunction/FiniteDimension.lean | 3 ++- .../Analysis/Calculus/ContDiff/Bounds.lean | 4 ++-- .../Calculus/ContDiff/FaaDiBruno.lean | 2 +- .../Analysis/Calculus/FDeriv/Bilinear.lean | 2 +- .../LineDeriv/IntegrationByParts.lean | 8 ++++---- Mathlib/Analysis/Calculus/MeanValue.lean | 2 +- Mathlib/Analysis/Complex/Isometry.lean | 2 +- .../Complex/UpperHalfPlane/Topology.lean | 2 +- Mathlib/Analysis/Convex/Birkhoff.lean | 4 ++-- Mathlib/Analysis/Convex/Combination.lean | 2 +- Mathlib/Analysis/Convex/Cone/Extension.lean | 2 +- Mathlib/Analysis/Convex/Jensen.lean | 2 +- .../Fourier/RiemannLebesgueLemma.lean | 2 +- Mathlib/Analysis/InnerProductSpace/PiL2.lean | 7 ++++--- .../InnerProductSpace/Projection.lean | 2 +- .../Normed/Group/ControlledClosure.lean | 4 ++-- Mathlib/Analysis/Normed/Operator/Banach.lean | 2 +- Mathlib/Analysis/NormedSpace/RieszLemma.lean | 2 +- .../SpecialFunctions/Complex/LogBounds.lean | 3 ++- .../ContinuousFunctionalCalculus/PosPart.lean | 18 +++++++++-------- .../Analysis/SpecialFunctions/Integrals.lean | 2 +- .../Abelian/EpiWithInjectiveKernel.lean | 2 +- Mathlib/CategoryTheory/Extensive.lean | 4 ++-- .../GradedObject/Trifunctor.lean | 4 ++-- .../Idempotents/FunctorCategories.lean | 4 ++-- .../LimitsOfProductsAndEqualizers.lean | 20 +++++++++---------- .../Limits/Constructions/Pullbacks.lean | 4 +++- .../FilteredColimitCommutesProduct.lean | 4 ++-- .../CategoryTheory/Limits/SmallComplete.lean | 4 ++-- Mathlib/CategoryTheory/Limits/VanKampen.lean | 13 ++++++------ .../Localization/CalculusOfFractions.lean | 14 ++++++------- Mathlib/CategoryTheory/Monoidal/Functor.lean | 2 +- Mathlib/CategoryTheory/Monoidal/Mon_.lean | 4 ++-- .../MorphismProperty/Limits.lean | 2 +- Mathlib/CategoryTheory/Quotient/Linear.lean | 12 +++++------ .../Sites/Coherent/Comparison.lean | 4 ++-- .../CategoryTheory/Sites/CoverLifting.lean | 2 +- .../CategoryTheory/Sites/CoverPreserving.lean | 2 +- .../Sites/DenseSubsite/SheafEquiv.lean | 5 +++-- .../Sites/EffectiveEpimorphic.lean | 10 +++++----- Mathlib/CategoryTheory/Sites/Over.lean | 2 +- .../CategoryTheory/Triangulated/Functor.lean | 2 +- Mathlib/Combinatorics/Additive/Energy.lean | 4 ++-- Mathlib/Combinatorics/Hall/Basic.lean | 9 +++++---- Mathlib/Combinatorics/Hall/Finite.lean | 3 ++- Mathlib/Data/Finite/Card.lean | 2 +- Mathlib/Data/Fintype/Card.lean | 2 +- Mathlib/Data/List/Indexes.lean | 5 ++--- Mathlib/Data/Real/Pi/Irrational.lean | 2 +- Mathlib/Data/Set/Card.lean | 4 ++-- Mathlib/Data/Setoid/Partition.lean | 2 +- Mathlib/FieldTheory/Extension.lean | 2 +- .../Euclidean/Angle/Oriented/Affine.lean | 4 ++-- .../Euclidean/Angle/Unoriented/Affine.lean | 4 ++-- Mathlib/Geometry/Euclidean/Circumcenter.lean | 2 +- Mathlib/Geometry/Euclidean/MongePoint.lean | 4 ++-- Mathlib/GroupTheory/Exponent.lean | 2 +- Mathlib/GroupTheory/Goursat.lean | 17 ++++++++-------- Mathlib/GroupTheory/HNNExtension.lean | 4 ++-- Mathlib/GroupTheory/PushoutI.lean | 11 +++++----- .../AffineSpace/Combination.lean | 10 +++++----- .../AffineSpace/Independent.lean | 6 +++--- Mathlib/LinearAlgebra/Charpoly/ToMatrix.lean | 3 ++- .../LinearAlgebra/CliffordAlgebra/Basic.lean | 2 +- .../LinearAlgebra/Dimension/DivisionRing.lean | 2 +- Mathlib/LinearAlgebra/Dual.lean | 3 ++- .../FreeModule/IdealQuotient.lean | 4 +--- Mathlib/LinearAlgebra/Matrix/Block.lean | 4 ++-- .../LinearAlgebra/Matrix/Charpoly/Coeff.lean | 6 +++--- .../LinearAlgebra/Matrix/Transvection.lean | 8 ++++---- Mathlib/LinearAlgebra/PID.lean | 2 +- Mathlib/LinearAlgebra/PerfectPairing.lean | 6 +++--- Mathlib/LinearAlgebra/Semisimple.lean | 2 +- Mathlib/Logic/Function/Basic.lean | 2 +- Mathlib/MeasureTheory/Constructions/Pi.lean | 7 ++++--- .../Constructions/Polish/Basic.lean | 4 ++-- .../MeasureTheory/Covering/Besicovitch.lean | 6 +++--- .../ConditionalExpectation/Unique.lean | 2 +- Mathlib/MeasureTheory/Function/LpSpace.lean | 2 +- .../Group/GeometryOfNumbers.lean | 2 +- Mathlib/MeasureTheory/Integral/Bochner.lean | 3 ++- Mathlib/MeasureTheory/Integral/Marginal.lean | 5 +++-- .../Integral/MeanInequalities.lean | 4 ++-- Mathlib/MeasureTheory/Integral/SetToL1.lean | 2 +- .../MeasurableSpace/Embedding.lean | 2 +- .../MeasureTheory/Measure/AEMeasurable.lean | 2 +- .../MeasureTheory/Measure/Haar/OfBasis.lean | 6 +++--- .../MeasureTheory/Measure/Haar/Unique.lean | 2 +- Mathlib/MeasureTheory/Measure/Prod.lean | 2 +- .../MeasureTheory/Measure/WithDensity.lean | 6 +++--- .../OuterMeasure/OfFunction.lean | 2 +- Mathlib/ModelTheory/Fraisse.lean | 2 +- .../NumberTheory/LSeries/HurwitzZetaOdd.lean | 2 +- .../LegendreSymbol/QuadraticChar/Basic.lean | 2 +- Mathlib/NumberTheory/MaricaSchoenheim.lean | 4 ++-- Mathlib/NumberTheory/Multiplicity.lean | 4 ++-- .../NumberTheory/NumberField/Embeddings.lean | 2 +- Mathlib/NumberTheory/Padics/Hensel.lean | 2 +- .../NumberTheory/Padics/PadicIntegers.lean | 4 ++-- Mathlib/NumberTheory/Padics/RingHoms.lean | 3 ++- Mathlib/NumberTheory/WellApproximable.lean | 7 ++++--- Mathlib/Order/Category/NonemptyFinLinOrd.lean | 2 +- Mathlib/Order/KrullDimension.lean | 4 ++-- Mathlib/Probability/StrongLaw.lean | 10 +++++----- Mathlib/RingTheory/Adjoin/PowerBasis.lean | 2 +- .../RankAndCardinality.lean | 4 ++-- Mathlib/RingTheory/DedekindDomain/PID.lean | 2 +- Mathlib/RingTheory/IsTensorProduct.lean | 4 ++-- .../RingTheory/Localization/BaseChange.lean | 4 ++-- Mathlib/RingTheory/Localization/Ideal.lean | 4 ++-- Mathlib/RingTheory/PowerSeries/Order.lean | 4 ++-- .../RingHom/FinitePresentation.lean | 2 +- Mathlib/SetTheory/Ordinal/Notation.lean | 2 +- .../Topology/Algebra/Group/OpenMapping.lean | 2 +- .../Topology/Algebra/Module/LinearPMap.lean | 2 +- .../Topology/Algebra/ProperAction/Basic.lean | 6 +++--- .../Topology/Category/Profinite/Nobeling.lean | 2 +- .../Category/TopCat/EffectiveEpi.lean | 3 ++- Mathlib/Topology/Connected/PathConnected.lean | 2 -- Mathlib/Topology/ContinuousMap/Algebra.lean | 4 ++-- .../ContinuousMap/StoneWeierstrass.lean | 8 ++++---- .../Topology/FiberBundle/Constructions.lean | 2 +- .../SheafCondition/PairwiseIntersections.lean | 2 +- Mathlib/Topology/UniformSpace/Completion.lean | 2 +- 149 files changed, 326 insertions(+), 303 deletions(-) diff --git a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean index 1d40a7cc808f9..f1d6988e3fced 100644 --- a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean +++ b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean @@ -645,7 +645,7 @@ def HomEquiv.fromExtendScalars {X Y} (g : X ⟶ (restrictScalars f).obj Y) : {toFun := fun s => HomEquiv.evalAt f s g, map_add' := fun (s₁ s₂ : S) => ?_, map_smul' := fun (r : R) (s : S) => ?_} · ext - dsimp only [evalAt_apply, LinearMap.add_apply] + dsimp only [m2, evalAt_apply, LinearMap.add_apply] rw [← add_smul] · ext x apply mul_smul (f r) s (g x) diff --git a/Mathlib/Algebra/Category/ModuleCat/Differentials/Basic.lean b/Mathlib/Algebra/Category/ModuleCat/Differentials/Basic.lean index cb4c1a9e98df5..a0ca4297cf81c 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Differentials/Basic.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Differentials/Basic.lean @@ -99,9 +99,9 @@ variable (f) in /-- The (universal) derivation in `(KaehlerDifferential f).Derivation f` when `f : A ⟶ B` is a morphism in the category `CommRingCat`. -/ noncomputable def D : (KaehlerDifferential f).Derivation f := - letI := f.toAlgebra + letI algebra := f.toAlgebra ModuleCat.Derivation.mk - (fun b ↦ _root_.KaehlerDifferential.D A B b) (by simp) (by simp) + (fun b ↦ _root_.KaehlerDifferential.D A B b) (by simp [algebra]) (by simp [algebra]) (_root_.KaehlerDifferential.D A B).map_algebraMap /-- When `f : A ⟶ B` is a morphism in the category `CommRingCat`, this is the diff --git a/Mathlib/Algebra/Homology/BifunctorAssociator.lean b/Mathlib/Algebra/Homology/BifunctorAssociator.lean index 3edcc3044b554..e7935213a0d90 100644 --- a/Mathlib/Algebra/Homology/BifunctorAssociator.lean +++ b/Mathlib/Algebra/Homology/BifunctorAssociator.lean @@ -341,8 +341,8 @@ lemma d_eq (j j' : ι₄) [HasGoodTrifunctor₁₂Obj F₁₂ G K₁ K₂ K₃ c by_cases h₁ : c₁₂.Rel i₁₂ (c₁₂.next i₁₂) · by_cases h₂ : ComplexShape.π c₁₂ c₃ c₄ (c₁₂.next i₁₂, i₃) = j' · rw [mapBifunctor.d₁_eq _ _ _ _ h₁ _ _ h₂] - simp only [mapBifunctor.d_eq, Functor.map_add, NatTrans.app_add, Preadditive.add_comp, - smul_add, Preadditive.comp_add, Linear.comp_units_smul] + simp only [i₁₂, mapBifunctor.d_eq, Functor.map_add, NatTrans.app_add, + Preadditive.add_comp, smul_add, Preadditive.comp_add, Linear.comp_units_smul] congr 1 · rw [← NatTrans.comp_app_assoc, ← Functor.map_comp, mapBifunctor.ι_D₁] @@ -701,7 +701,8 @@ lemma d_eq : · rw [d₃_eq_zero] intro h₂ apply h₁ - simpa only [ComplexShape.next_π₂ c₂ c₂₃ i₂ h₂] using ComplexShape.rel_π₂ c₂ c₂₃ i₂ h₂ + simpa only [i₂₃, ComplexShape.next_π₂ c₂ c₂₃ i₂ h₂] + using ComplexShape.rel_π₂ c₂ c₂₃ i₂ h₂ end mapBifunctor₂₃ diff --git a/Mathlib/Algebra/Homology/DerivedCategory/Ext/Basic.lean b/Mathlib/Algebra/Homology/DerivedCategory/Ext/Basic.lean index b5130066b0f40..7a3a1921a4ba1 100644 --- a/Mathlib/Algebra/Homology/DerivedCategory/Ext/Basic.lean +++ b/Mathlib/Algebra/Homology/DerivedCategory/Ext/Basic.lean @@ -206,33 +206,33 @@ private lemma zero_hom' : (0 : Ext X Y n).hom' = 0 := @[simp] lemma add_comp (α₁ α₂ : Ext X Y n) {m : ℕ} (β : Ext Y Z m) {p : ℕ} (h : n + m = p) : (α₁ + α₂).comp β h = α₁.comp β h + α₂.comp β h := by - letI := HasDerivedCategory.standard C; ext; simp [add_hom'] + letI := HasDerivedCategory.standard C; ext; simp [this, add_hom'] @[simp] lemma comp_add (α : Ext X Y n) {m : ℕ} (β₁ β₂ : Ext Y Z m) {p : ℕ} (h : n + m = p) : α.comp (β₁ + β₂) h = α.comp β₁ h + α.comp β₂ h := by - letI := HasDerivedCategory.standard C; ext; simp [add_hom'] + letI := HasDerivedCategory.standard C; ext; simp [this, add_hom'] @[simp] lemma neg_comp (α : Ext X Y n) {m : ℕ} (β : Ext Y Z m) {p : ℕ} (h : n + m = p) : (-α).comp β h = -α.comp β h := by - letI := HasDerivedCategory.standard C; ext; simp [neg_hom'] + letI := HasDerivedCategory.standard C; ext; simp [this, neg_hom'] @[simp] lemma comp_neg (α : Ext X Y n) {m : ℕ} (β : Ext Y Z m) {p : ℕ} (h : n + m = p) : α.comp (-β) h = -α.comp β h := by - letI := HasDerivedCategory.standard C; ext; simp [neg_hom'] + letI := HasDerivedCategory.standard C; ext; simp [this, neg_hom'] variable (X n) in @[simp] lemma zero_comp {m : ℕ} (β : Ext Y Z m) (p : ℕ) (h : n + m = p) : (0 : Ext X Y n).comp β h = 0 := by - letI := HasDerivedCategory.standard C; ext; simp [zero_hom'] + letI := HasDerivedCategory.standard C; ext; simp [this, zero_hom'] @[simp] lemma comp_zero (α : Ext X Y n) (Z : C) (m : ℕ) (p : ℕ) (h : n + m = p) : α.comp (0 : Ext Y Z m) h = 0 := by - letI := HasDerivedCategory.standard C; ext; simp [zero_hom'] + letI := HasDerivedCategory.standard C; ext; simp [this, zero_hom'] @[simp] lemma mk₀_id_comp (α : Ext X Y n) : @@ -247,7 +247,7 @@ lemma comp_mk₀_id (α : Ext X Y n) : variable (X Y) in @[simp] lemma mk₀_zero : mk₀ (0 : X ⟶ Y) = 0 := by - letI := HasDerivedCategory.standard C; ext; simp [zero_hom'] + letI := HasDerivedCategory.standard C; ext; simp [this, zero_hom'] section diff --git a/Mathlib/Algebra/Homology/ShortComplex/Abelian.lean b/Mathlib/Algebra/Homology/ShortComplex/Abelian.lean index 7bc3d040b42d2..ad19544a15033 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/Abelian.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/Abelian.lean @@ -84,7 +84,8 @@ noncomputable def ofAbelian : S.LeftHomologyData := by IsLimit.conePointUniqueUpToIso_hom_comp _ _ WalkingParallelPair.zero have fac : f' = Abelian.factorThruImage S.f ≫ e.hom ≫ kernel.ι γ := by rw [hf', he] - simp only [f', kernel.lift_ι, abelianImageToKernel, ← cancel_mono (kernel.ι S.g), assoc] + simp only [γ, f', kernel.lift_ι, abelianImageToKernel, ← cancel_mono (kernel.ι S.g), + assoc] have hπ : IsColimit (CokernelCofork.ofπ _ wπ) := CokernelCofork.IsColimit.ofπ _ _ (fun x hx => cokernel.desc _ x (by @@ -150,7 +151,7 @@ noncomputable def ofAbelian : S.RightHomologyData := by IsColimit.comp_coconePointUniqueUpToIso_hom _ _ WalkingParallelPair.one have fac : g' = cokernel.π γ ≫ e.hom ≫ Abelian.factorThruCoimage S.g := by rw [hg', reassoc_of% he] - simp only [g', cokernel.π_desc, ← cancel_epi (cokernel.π S.f), + simp only [γ, g', cokernel.π_desc, ← cancel_epi (cokernel.π S.f), cokernel_π_comp_cokernelToAbelianCoimage_assoc] have hι : IsLimit (KernelFork.ofι _ wι) := KernelFork.IsLimit.ofι _ _ diff --git a/Mathlib/Algebra/Lie/CartanExists.lean b/Mathlib/Algebra/Lie/CartanExists.lean index 5ee149734e3b5..cca8926c878a6 100644 --- a/Mathlib/Algebra/Lie/CartanExists.lean +++ b/Mathlib/Algebra/Lie/CartanExists.lean @@ -320,7 +320,7 @@ lemma engel_isBot_of_isMin (hLK : finrank K L ≤ #K) (U : LieSubalgebra K L) rw [← hn] clear hn induction n with - | zero => simp only [pow_zero, LinearMap.one_apply] + | zero => simp only [z', pow_zero, LinearMap.one_apply] | succ n ih => rw [pow_succ', pow_succ', LinearMap.mul_apply, ih]; rfl classical -- Now let `n` be the smallest power such that `⁅v, _⁆ ^ n` kills `z'`. diff --git a/Mathlib/Algebra/Lie/Weights/Chain.lean b/Mathlib/Algebra/Lie/Weights/Chain.lean index 01285cf51e542..cc0efd7803997 100644 --- a/Mathlib/Algebra/Lie/Weights/Chain.lean +++ b/Mathlib/Algebra/Lie/Weights/Chain.lean @@ -213,19 +213,19 @@ lemma exists_forall_mem_corootSpace_smul_add_eq_zero exact SetCoe.ext <| smul_left_injective ℤ hα <| by rwa [add_left_inj] at hij have h₂ : ∀ i, MapsTo (toEnd R H M x) ↑(N i) ↑(N i) := fun _ _ ↦ LieSubmodule.lie_mem _ have h₃ : genWeightSpaceChain M α χ p q = ⨆ i ∈ Finset.Ioo p q, N i := by - simp_rw [genWeightSpaceChain_def', LieSubmodule.iSup_coe_toSubmodule] + simp_rw [N, genWeightSpaceChain_def', LieSubmodule.iSup_coe_toSubmodule] rw [← trace_toEnd_genWeightSpaceChain_eq_zero M α χ p q hp hq hx, ← LieSubmodule.toEnd_restrict_eq_toEnd] -- The lines below illustrate the cost of treating `LieSubmodule` as both a -- `Submodule` and a `LieSubmodule` simultaneously. erw [LinearMap.trace_eq_sum_trace_restrict_of_eq_biSup _ h₁ h₂ (genWeightSpaceChain M α χ p q) h₃] - simp_rw [LieSubmodule.toEnd_restrict_eq_toEnd] + simp_rw [N, LieSubmodule.toEnd_restrict_eq_toEnd] dsimp [N] convert_to _ = ∑ k ∈ Finset.Ioo p q, (LinearMap.trace R { x // x ∈ (genWeightSpace M (k • α + χ)) }) ((toEnd R { x // x ∈ H } { x // x ∈ genWeightSpace M (k • α + χ) }) x) - simp_rw [trace_toEnd_genWeightSpace, Pi.add_apply, Pi.smul_apply, smul_add, ← smul_assoc, - Finset.sum_add_distrib, ← Finset.sum_smul, natCast_zsmul] + simp_rw [a, b, trace_toEnd_genWeightSpace, Pi.add_apply, Pi.smul_apply, smul_add, + ← smul_assoc, Finset.sum_add_distrib, ← Finset.sum_smul, natCast_zsmul] end IsCartanSubalgebra diff --git a/Mathlib/Algebra/Lie/Weights/Killing.lean b/Mathlib/Algebra/Lie/Weights/Killing.lean index 7c2b6ece5a6b5..a4667d0fba6aa 100644 --- a/Mathlib/Algebra/Lie/Weights/Killing.lean +++ b/Mathlib/Algebra/Lie/Weights/Killing.lean @@ -298,7 +298,7 @@ lemma isSemisimple_ad_of_mem_isCartanSubalgebra {x : L} (hx : x ∈ H) : /- `y` commutes with all elements of `H` because `S` has eigenvalue 0 on `H`, `S = ad K L y`. -/ have hy' (z : L) (hz : z ∈ H) : ⁅y, z⁆ = 0 := by rw [← LieSubalgebra.mem_toLieSubmodule, ← rootSpace_zero_eq] at hz - simp [← ad_apply (R := K), ← LieDerivation.coe_ad_apply_eq_ad_apply, hy, aux hz] + simp [S', ← ad_apply (R := K), ← LieDerivation.coe_ad_apply_eq_ad_apply, hy, aux hz] /- Thus `y` belongs to `H` since `H` is self-normalizing. -/ replace hy' : y ∈ H := by suffices y ∈ H.normalizer by rwa [LieSubalgebra.IsCartanSubalgebra.self_normalizing] at this @@ -489,7 +489,7 @@ lemma exists_isSl2Triple_of_weight_isNonZero {α : Weight K H L} (hα : α.IsNon let h : H := ⟨⁅e, f'⁆, hef ▸ Submodule.smul_mem _ _ (Submodule.coe_mem _)⟩ have hh : α h ≠ 0 := by have : h = killingForm K L e f' • (cartanEquivDual H).symm α := by - simp only [Subtype.ext_iff, hef] + simp only [h, Subtype.ext_iff, hef] rw [Submodule.coe_smul_of_tower] rw [this, map_smul, smul_eq_mul, ne_eq, mul_eq_zero, not_or] exact ⟨hf, root_apply_cartanEquivDual_symm_ne_zero hα⟩ diff --git a/Mathlib/Algebra/Module/FinitePresentation.lean b/Mathlib/Algebra/Module/FinitePresentation.lean index 8d27146b99787..468b2cc116e18 100644 --- a/Mathlib/Algebra/Module/FinitePresentation.lean +++ b/Mathlib/Algebra/Module/FinitePresentation.lean @@ -438,7 +438,7 @@ lemma Module.FinitePresentation.exists_lift_equiv_of_isLocalizedModule ext x apply ((Module.End_isUnit_iff _).mp (IsLocalizedModule.map_units g s)).1 have : ∀ x, g (l' x) = s.1 • (l (f x)) := LinearMap.congr_fun H - simp only [LinearMap.coe_comp, LinearMap.coe_restrictScalars, LinearEquiv.coe_coe, + simp only [rs, LinearMap.coe_comp, LinearMap.coe_restrictScalars, LinearEquiv.coe_coe, Function.comp_apply, LocalizedModule.mkLinearMap_apply, LinearEquiv.ofBijective_apply, LinearMap.smul_apply, LocalizedModule.map_mk, algebraMap_end_apply] rw [← map_smul, ← smul_assoc, Algebra.smul_def s.1, hsu.mul_val_inv, one_smul] diff --git a/Mathlib/Algebra/Module/FreeLocus.lean b/Mathlib/Algebra/Module/FreeLocus.lean index 68a8d2d46a780..770e2b8cab031 100644 --- a/Mathlib/Algebra/Module/FreeLocus.lean +++ b/Mathlib/Algebra/Module/FreeLocus.lean @@ -89,10 +89,10 @@ lemma comap_freeLocus_le {A} [CommRing A] [Algebra R A] : let Aₚ := Localization.AtPrime p.asIdeal rw [Set.mem_preimage, mem_freeLocus_iff_tensor _ Rₚ] at hp rw [mem_freeLocus_iff_tensor _ Aₚ] - letI : Algebra Rₚ Aₚ := (Localization.localRingHom + letI algebra : Algebra Rₚ Aₚ := (Localization.localRingHom (comap (algebraMap R A) p).asIdeal p.asIdeal (algebraMap R A) rfl).toAlgebra have : IsScalarTower R Rₚ Aₚ := IsScalarTower.of_algebraMap_eq' - (by simp [RingHom.algebraMap_toAlgebra, Localization.localRingHom, + (by simp [Rₚ, Aₚ, algebra, RingHom.algebraMap_toAlgebra, Localization.localRingHom, ← IsScalarTower.algebraMap_eq]) let e := AlgebraTensorModule.cancelBaseChange R Rₚ Aₚ Aₚ M ≪≫ₗ (AlgebraTensorModule.cancelBaseChange R A Aₚ Aₚ M).symm @@ -125,7 +125,7 @@ lemma freeLocus_localization (S : Submonoid R) : letI : Module (Localization S) Mₚ := Module.compHom Mₚ (algebraMap _ Rₚ) have : IsScalarTower R (Localization S) Mₚ := ⟨fun r r' m ↦ show algebraMap _ Rₚ (r • r') • m = _ by - simp [Algebra.smul_def, ← IsScalarTower.algebraMap_apply, mul_smul]; rfl⟩ + simp [p', Rₚ, Mₚ, Algebra.smul_def, ← IsScalarTower.algebraMap_apply, mul_smul]; rfl⟩ have : IsScalarTower (Localization S) Rₚ Mₚ := ⟨fun r r' m ↦ show _ = algebraMap _ Rₚ r • r' • m by rw [← mul_smul, ← Algebra.smul_def]⟩ let l := (IsLocalizedModule.liftOfLE _ _ hp' (LocalizedModule.mkLinearMap S M) @@ -202,7 +202,7 @@ lemma isLocallyConstant_rankAtStalk_freeLocus [Module.FinitePresentation R M] : letI : Module (Localization.Away f) Mₚ := Module.compHom Mₚ (algebraMap _ Rₚ) have : IsScalarTower R (Localization.Away f) Mₚ := ⟨fun r r' m ↦ show algebraMap _ Rₚ (r • r') • m = _ by - simp [Algebra.smul_def, ← IsScalarTower.algebraMap_apply, mul_smul]; rfl⟩ + simp [Rₚ, Mₚ, Algebra.smul_def, ← IsScalarTower.algebraMap_apply, mul_smul]; rfl⟩ have : IsScalarTower (Localization.Away f) Rₚ Mₚ := ⟨fun r r' m ↦ show _ = algebraMap _ Rₚ r • r' • m by rw [← mul_smul, ← Algebra.smul_def]⟩ let l := (IsLocalizedModule.liftOfLE _ _ hp' (LocalizedModule.mkLinearMap (.powers f) M) diff --git a/Mathlib/Algebra/Module/LinearMap/Polynomial.lean b/Mathlib/Algebra/Module/LinearMap/Polynomial.lean index 3a315ebcd1b5c..5637e7a41bdf3 100644 --- a/Mathlib/Algebra/Module/LinearMap/Polynomial.lean +++ b/Mathlib/Algebra/Module/LinearMap/Polynomial.lean @@ -548,12 +548,12 @@ lemma exists_isNilRegular_of_finrank_le_card (h : finrank R M ≤ #R) : have aux : ((polyCharpoly φ b).coeff (nilRank φ)).IsHomogeneous (n - nilRank φ) := polyCharpoly_coeff_isHomogeneous _ b (nilRank φ) (n - nilRank φ) - (by simp [nilRank_le_card φ bₘ, finrank_eq_card_chooseBasisIndex]) + (by simp [n, nilRank_le_card φ bₘ, finrank_eq_card_chooseBasisIndex]) obtain ⟨x, hx⟩ : ∃ r, eval r ((polyCharpoly _ b).coeff (nilRank φ)) ≠ 0 := by by_contra! h₀ apply polyCharpoly_coeff_nilRank_ne_zero φ b apply aux.eq_zero_of_forall_eval_eq_zero_of_le_card h₀ (le_trans _ h) - simp only [finrank_eq_card_chooseBasisIndex, Nat.cast_le, Nat.sub_le] + simp only [n, finrank_eq_card_chooseBasisIndex, Nat.cast_le, Nat.sub_le] let c := Finsupp.equivFunOnFinite.symm x use b.repr.symm c rwa [isNilRegular_iff_coeff_polyCharpoly_nilRank_ne_zero _ b, LinearEquiv.apply_symm_apply] diff --git a/Mathlib/Algebra/Order/Sub/WithTop.lean b/Mathlib/Algebra/Order/Sub/WithTop.lean index f6b0e29c1f838..5b2d1da2ff61f 100644 --- a/Mathlib/Algebra/Order/Sub/WithTop.lean +++ b/Mathlib/Algebra/Order/Sub/WithTop.lean @@ -76,6 +76,7 @@ instance : OrderedSub (WithTop α) := by · simp cases z · simp - norm_cast; exact tsub_le_iff_right + norm_cast + exact tsub_le_iff_right end WithTop diff --git a/Mathlib/Algebra/Polynomial/Div.lean b/Mathlib/Algebra/Polynomial/Div.lean index 41009d47c194c..41487703c5334 100644 --- a/Mathlib/Algebra/Polynomial/Div.lean +++ b/Mathlib/Algebra/Polynomial/Div.lean @@ -434,7 +434,7 @@ lemma coeff_divByMonic_X_sub_C_rec (p : R[X]) (a : R) (n : ℕ) : rw [← p.modByMonic_add_div this] have : degree (p %ₘ (X - C a)) < ↑(n + 1) := degree_X_sub_C a ▸ p.degree_modByMonic_lt this |>.trans_le <| WithBot.coe_le_coe.mpr le_add_self - simp [sub_mul, add_sub, coeff_eq_zero_of_degree_lt this] + simp [q, sub_mul, add_sub, coeff_eq_zero_of_degree_lt this] theorem coeff_divByMonic_X_sub_C (p : R[X]) (a : R) (n : ℕ) : (p /ₘ (X - C a)).coeff n = ∑ i ∈ Icc (n + 1) p.natDegree, a ^ (i - (n + 1)) * p.coeff i := by diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean index 251acb43cbb94..faa80b6bd843f 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean @@ -326,13 +326,13 @@ lemma XYIdeal_neg_mul {x y : F} (h : W.Nonsingular x y) : refine ⟨C <| C W_X⁻¹ * -(X + C (2 * x + W.a₂)), C <| C <| W_X⁻¹ * W.a₁, 0, C <| C <| W_X⁻¹ * -1, ?_⟩ rw [← mul_right_inj' <| C_ne_zero.mpr <| C_ne_zero.mpr hx] - simp only [mul_add, ← mul_assoc, ← C_mul, mul_inv_cancel₀ hx] + simp only [W_X, mul_add, ← mul_assoc, ← C_mul, mul_inv_cancel₀ hx] C_simp ring1 · let W_Y := 2 * y + W.a₁ * x + W.a₃ refine ⟨0, C <| C W_Y⁻¹, C <| C <| W_Y⁻¹ * -1, 0, ?_⟩ rw [negY, ← mul_right_inj' <| C_ne_zero.mpr <| C_ne_zero.mpr hy] - simp only [mul_add, ← mul_assoc, ← C_mul, mul_inv_cancel₀ hy] + simp only [W_Y, mul_add, ← mul_assoc, ← C_mul, mul_inv_cancel₀ hy] C_simp ring1 @@ -374,7 +374,7 @@ lemma XYIdeal_mul_XYIdeal {x₁ x₂ y₁ y₂ : F} (h₁ : W.Equation x₁ y₁ 0, C (C y⁻¹) * (Y - W.negPolynomial), ?_⟩, by rw [map_add, map_one, _root_.map_mul <| mk W, AdjoinRoot.mk_self, mul_zero, add_zero]⟩ rw [polynomial, negPolynomial, ← mul_right_inj' <| C_ne_zero.mpr <| C_ne_zero.mpr hxy] - simp only [mul_add, ← mul_assoc, ← C_mul, mul_inv_cancel₀ hxy] + simp only [y, mul_add, ← mul_assoc, ← C_mul, mul_inv_cancel₀ hxy] linear_combination (norm := (rw [b₂, b₄, negY]; C_simp; ring1)) -4 * congr_arg C (congr_arg C <| (equation_iff ..).mp h₁) · replace hx := sub_ne_zero_of_ne hx diff --git a/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean b/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean index 66aff1bb5a74d..31a833ebabe90 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean @@ -113,10 +113,12 @@ lemma isEmbedding_pullback {X Y S : Scheme.{u}} (f : X ⟶ S) (g : Y ⟶ S) [Sur obtain ⟨x, rfl⟩ := (Scheme.homeoOfIso (pullbackSpecIso R A B).symm).surjective x simp only [Scheme.homeoOfIso_apply, Function.comp_apply] ext - · simp only [← Scheme.comp_base_apply, pullback.lift_fst, Iso.symm_hom, Iso.inv_hom_id] + · simp only [L, ← Scheme.comp_base_apply, pullback.lift_fst, Iso.symm_hom, + Iso.inv_hom_id] erw [← Scheme.comp_base_apply, pullbackSpecIso_inv_fst_assoc] rfl - · simp only [← Scheme.comp_base_apply, pullback.lift_snd, Iso.symm_hom, Iso.inv_hom_id] + · simp only [L, ← Scheme.comp_base_apply, pullback.lift_snd, Iso.symm_hom, + Iso.inv_hom_id] erw [← Scheme.comp_base_apply, pullbackSpecIso_inv_snd_assoc] rfl let 𝒰 := S.affineOpenCover.openCover @@ -178,7 +180,8 @@ lemma isEmbedding_pullback {X Y S : Scheme.{u}} (f : X ⟶ S) (g : Y ⟶ S) [Sur inferInstance inferInstance inferInstance convert this using 6 apply pullback.hom_ext <;> - simp [𝓤, ← pullback.condition, ← pullback.condition_assoc, Scheme.Cover.pullbackHom] + simp [𝓤, ← pullback.condition, ← pullback.condition_assoc, + Scheme.Cover.pullbackHom] end SurjectiveOnStalks diff --git a/Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean b/Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean index 84a2cb83d773a..d8cdaa1f4dd0a 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean @@ -135,7 +135,7 @@ lemma compactSpace_of_universallyClosed contrapose! h obtain ⟨x, hx⟩ := h obtain ⟨z, rfl, hzr⟩ := exists_preimage_pullback x t' (Subsingleton.elim (f.base x) (q.base t')) - suffices ∀ i, t ∈ (Ti i).comap (comap φ) → p.base z ∉ U i from ⟨z, by simpa [Z, p, hzr], hzr⟩ + suffices ∀ i, t ∈ (Ti i).comap (comap φ) → p.base z ∉ U i from ⟨z, by simpa [Z, p, fT, hzr], hzr⟩ intro i hi₁ hi₂ rw [comap_basicOpen, show φ (.X i) = 0 by simpa [φ] using (hx i · hi₂), basicOpen_zero] at hi₁ cases hi₁ diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index 96a4c71a98a96..a209d03a31d28 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -565,13 +565,13 @@ theorem toPiLocalization_surjective_of_discreteTopology : let e := Equiv.ofInjective f fun p q eq ↦ Set.singleton_injective (hf p ▸ eq ▸ hf q) have loc a : IsLocalization.AtPrime (Localization.Away a.1) (e.symm a).1 := (isLocalization_away_iff_atPrime_of_basicOpen_eq_singleton <| hf _).mp <| by - simp_rw [Equiv.apply_ofInjective_symm]; infer_instance + simp_rw [e, Equiv.apply_ofInjective_symm]; infer_instance let algE a := IsLocalization.algEquiv (e.symm a).1.primeCompl (Localization.AtPrime (e.symm a).1) (Localization.Away a.1) have span_eq : Ideal.span (Set.range f) = ⊤ := iSup_basicOpen_eq_top_iff.mp <| top_unique fun p _ ↦ TopologicalSpace.Opens.mem_iSup.mpr ⟨p, (hf p).ge rfl⟩ replace hf a : (basicOpen a.1 : Set _) = {e.symm a} := by - simp_rw [← hf, Equiv.apply_ofInjective_symm] + simp_rw [e, ← hf, Equiv.apply_ofInjective_symm] obtain ⟨r, eq, -⟩ := Localization.existsUnique_algebraMap_eq_of_span_eq_top _ span_eq (fun a ↦ algE a (x _)) fun a b ↦ by obtain rfl | ne := eq_or_ne a b; · rfl diff --git a/Mathlib/AlgebraicGeometry/RationalMap.lean b/Mathlib/AlgebraicGeometry/RationalMap.lean index e9894ddfbcd61..7e848dcf6666c 100644 --- a/Mathlib/AlgebraicGeometry/RationalMap.lean +++ b/Mathlib/AlgebraicGeometry/RationalMap.lean @@ -511,8 +511,8 @@ def RationalMap.toPartialMap [IsReduced X] [Y.IsSeparated] (f : X ⤏ Y) : X.Par show _ ≫ _ ≫ (g x).hom = _ ≫ _ ≫ (g y).hom simp_rw [← cancel_epi (X.isoOfEq congr($(hg₂ x) ⊓ $(hg₂ y))).hom, ← Category.assoc] convert (PartialMap.equiv_iff_of_isSeparated (S := ⊤_ _) (f := g x) (g := g y)).mp ?_ using 1 - · dsimp; congr 1; simp [← cancel_mono (Opens.ι _)] - · dsimp; congr 1; simp [← cancel_mono (Opens.ι _)] + · dsimp; congr 1; simp [g, ← cancel_mono (Opens.ι _)] + · dsimp; congr 1; simp [g, ← cancel_mono (Opens.ι _)] · rw [← PartialMap.toRationalMap_eq_iff, hg₁, hg₁] lemma PartialMap.toPartialMap_toRationalMap_restrict [IsReduced X] [Y.IsSeparated] diff --git a/Mathlib/AlgebraicGeometry/Sites/MorphismProperty.lean b/Mathlib/AlgebraicGeometry/Sites/MorphismProperty.lean index e8684970f0fd7..84d50d0d2a13d 100644 --- a/Mathlib/AlgebraicGeometry/Sites/MorphismProperty.lean +++ b/Mathlib/AlgebraicGeometry/Sites/MorphismProperty.lean @@ -130,7 +130,7 @@ lemma grothendieckTopology_cover {X : Scheme.{u}} (𝒰 : Cover.{v} P X) : } refine ⟨_, pretopology_cover 𝒱, ?_⟩ rintro _ _ ⟨y⟩ - exact ⟨_, 𝟙 _, 𝒰.map (𝒰.f y), ⟨_⟩, by simp⟩ + exact ⟨_, 𝟙 _, 𝒰.map (𝒰.f y), ⟨_⟩, by simp [𝒱]⟩ section diff --git a/Mathlib/AlgebraicGeometry/ValuativeCriterion.lean b/Mathlib/AlgebraicGeometry/ValuativeCriterion.lean index dbc4ac5a16007..10d725b7f0215 100644 --- a/Mathlib/AlgebraicGeometry/ValuativeCriterion.lean +++ b/Mathlib/AlgebraicGeometry/ValuativeCriterion.lean @@ -242,9 +242,9 @@ lemma IsSeparated.of_valuativeCriterion [QuasiSeparated f] let S' : ValuativeCommSq f := ⟨S.R, S.K, S.i₁, S.i₂ ≫ pullback.fst f f ≫ f, hc⟩ have : Subsingleton S'.commSq.LiftStruct := hf S' let S'l₁ : S'.commSq.LiftStruct := ⟨S.i₂ ≫ pullback.fst f f, - by simp [← S.commSq.w_assoc], by simp⟩ + by simp [S', ← S.commSq.w_assoc], by simp [S']⟩ let S'l₂ : S'.commSq.LiftStruct := ⟨S.i₂ ≫ pullback.snd f f, - by simp [← S.commSq.w_assoc], by simp [pullback.condition]⟩ + by simp [S', ← S.commSq.w_assoc], by simp [S', pullback.condition]⟩ have h₁₂ : S'l₁ = S'l₂ := Subsingleton.elim _ _ constructor constructor diff --git a/Mathlib/AlgebraicTopology/AlternatingFaceMapComplex.lean b/Mathlib/AlgebraicTopology/AlternatingFaceMapComplex.lean index 25c2c3588b0e9..7e25833d9f89f 100644 --- a/Mathlib/AlgebraicTopology/AlternatingFaceMapComplex.lean +++ b/Mathlib/AlgebraicTopology/AlternatingFaceMapComplex.lean @@ -80,14 +80,14 @@ theorem d_squared (n : ℕ) : objD X (n + 1) ≫ objD X n = 0 := by apply Finset.sum_bij φ · -- φ(S) is contained in Sᶜ intro ij hij - simp only [S, Finset.mem_univ, Finset.compl_filter, Finset.mem_filter, true_and, + simp only [S, φ, Finset.mem_univ, Finset.compl_filter, Finset.mem_filter, true_and, Fin.val_succ, Fin.coe_castLT] at hij ⊢ linarith · -- φ : S → Sᶜ is injective rintro ⟨i, j⟩ hij ⟨i', j'⟩ hij' h rw [Prod.mk.inj_iff] - exact ⟨by simpa using congr_arg Prod.snd h, - by simpa [Fin.castSucc_castLT] using congr_arg Fin.castSucc (congr_arg Prod.fst h)⟩ + exact ⟨by simpa [φ] using congr_arg Prod.snd h, + by simpa [φ, Fin.castSucc_castLT] using congr_arg Fin.castSucc (congr_arg Prod.fst h)⟩ · -- φ : S → Sᶜ is surjective rintro ⟨i', j'⟩ hij' simp only [S, Finset.mem_univ, forall_true_left, Prod.forall, Finset.compl_filter, @@ -103,7 +103,7 @@ theorem d_squared (n : ℕ) : objD X (n + 1) ≫ objD X n = 0 := by dsimp simp only [zsmul_comp, comp_zsmul, smul_smul, ← neg_smul] congr 1 - · simp only [Fin.val_succ, pow_add, pow_one, mul_neg, neg_neg, mul_one] + · simp only [φ, Fin.val_succ, pow_add, pow_one, mul_neg, neg_neg, mul_one] apply mul_comm · rw [CategoryTheory.SimplicialObject.δ_comp_δ''] simpa [S] using hij diff --git a/Mathlib/AlgebraicTopology/DoldKan/Decomposition.lean b/Mathlib/AlgebraicTopology/DoldKan/Decomposition.lean index 7556fb09166e3..d1b827f3f5c7c 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/Decomposition.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/Decomposition.lean @@ -66,7 +66,7 @@ theorem decomposition_Q (n q : ℕ) : symm conv_rhs => rw [sub_eq_add_neg, add_comm] let q' : Fin (n + 1) := ⟨q, Nat.succ_le_iff.mp hqn⟩ - rw [← @Finset.add_sum_erase _ _ _ _ _ _ q' (by simp)] + rw [← @Finset.add_sum_erase _ _ _ _ _ _ q' (by simp [q'])] congr · have hnaq' : n = a + q := by omega simp only [Fin.val_mk, (HigherFacesVanish.of_P q n).comp_Hσ_eq hnaq', diff --git a/Mathlib/Analysis/Analytic/Inverse.lean b/Mathlib/Analysis/Analytic/Inverse.lean index 6efcaa447e727..6ef7fff93cce3 100644 --- a/Mathlib/Analysis/Analytic/Inverse.lean +++ b/Mathlib/Analysis/Analytic/Inverse.lean @@ -435,7 +435,7 @@ theorem radius_rightInv_pos_of_radius_pos_aux2 {x : E} {n : ℕ} (hn : 2 ≤ n + calc ∑ k ∈ Ico 1 (n + 1), a ^ k * ‖p.rightInv i x k‖ = a * I + ∑ k ∈ Ico 2 (n + 1), a ^ k * ‖p.rightInv i x k‖ := by - simp only [LinearIsometryEquiv.norm_map, pow_one, rightInv_coeff_one, + simp only [I, LinearIsometryEquiv.norm_map, pow_one, rightInv_coeff_one, show Ico (1 : ℕ) 2 = {1} from Nat.Ico_succ_singleton 1, sum_singleton, ← sum_Ico_consecutive _ one_le_two hn] _ = @@ -465,8 +465,8 @@ theorem radius_rightInv_pos_of_radius_pos_aux2 {x : E} {n : ℕ} (hn : 2 ≤ n + _ = I * a + I * C * ∑ k ∈ Ico 2 (n + 1), a ^ k * ∑ c ∈ ({c | 1 < Composition.length c}.toFinset : Finset (Composition k)), r ^ c.length * ∏ j, ‖p.rightInv i x (c.blocksFun j)‖ := by - simp_rw [mul_assoc C, ← mul_sum, ← mul_assoc, mul_comm _ ‖(i.symm : F →L[𝕜] E)‖, mul_assoc, - ← mul_sum, ← mul_assoc, mul_comm _ C, mul_assoc, ← mul_sum] + simp_rw [I, mul_assoc C, ← mul_sum, ← mul_assoc, mul_comm _ ‖(i.symm : F →L[𝕜] E)‖, + mul_assoc, ← mul_sum, ← mul_assoc, mul_comm _ C, mul_assoc, ← mul_sum] ring _ ≤ I * a + I * C * ∑ k ∈ Ico 2 (n + 1), (r * ∑ j ∈ Ico 1 n, a ^ j * ‖p.rightInv i x j‖) ^ k := by diff --git a/Mathlib/Analysis/BoxIntegral/UnitPartition.lean b/Mathlib/Analysis/BoxIntegral/UnitPartition.lean index 9c0e70ce75075..17d078e1e0df9 100644 --- a/Mathlib/Analysis/BoxIntegral/UnitPartition.lean +++ b/Mathlib/Analysis/BoxIntegral/UnitPartition.lean @@ -73,13 +73,14 @@ theorem BoxIntegral.le_hasIntegralVertices_of_isBounded [Finite ι] {s : Set (ι let C : ℕ := ⌈R⌉₊ have hC := Nat.ceil_pos.mpr hR₁ let I : Box ι := Box.mk (fun _ ↦ - C) (fun _ ↦ C ) - (fun _ ↦ by simp only [neg_lt_self_iff, Nat.cast_pos, hC]) + (fun _ ↦ by simp [C, neg_lt_self_iff, Nat.cast_pos, hC]) refine ⟨I, ⟨fun _ ↦ - C, fun _ ↦ C, fun i ↦ (Int.cast_neg_natCast C).symm, fun _ ↦ rfl⟩, le_trans hR₂ ?_⟩ suffices Metric.ball (0 : ι → ℝ) C ≤ I from le_trans (Metric.ball_subset_ball (Nat.le_ceil R)) this intro x hx - simp_rw [mem_ball_zero_iff, pi_norm_lt_iff (Nat.cast_pos.mpr hC), Real.norm_eq_abs, abs_lt] at hx + simp_rw [C, mem_ball_zero_iff, pi_norm_lt_iff (Nat.cast_pos.mpr hC), + Real.norm_eq_abs, abs_lt] at hx exact fun i ↦ ⟨(hx i).1, le_of_lt (hx i).2⟩ end hasIntegralVertices diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean index 76c6e66606ed2..a05dd23178aaa 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean @@ -465,7 +465,7 @@ lemma spectrum_star_mul_self_nonneg {b : A} : ∀ x ∈ spectrum ℝ (star b * b have h_c_spec₀ : SpectrumRestricts (- (star c * c)) (ContinuousMap.realToNNReal ·) := by simp only [SpectrumRestricts.nnreal_iff, h_eq_a_neg] rw [← cfc_pow _ _ (ha := .star_mul_self b)] - simp only [cfc_map_spectrum (R := ℝ) (fun x => (-ContinuousMap.id ℝ ⊔ 0) x ^ 3) (star b * b)] + simp only [a, cfc_map_spectrum (R := ℝ) (fun x => (-ContinuousMap.id ℝ ⊔ 0) x ^ 3) (star b * b)] rintro - ⟨x, -, rfl⟩ simp have c_eq := star_mul_self_add_self_mul_star c diff --git a/Mathlib/Analysis/Calculus/BumpFunction/FiniteDimension.lean b/Mathlib/Analysis/Calculus/BumpFunction/FiniteDimension.lean index 2f3e0be74bdfb..70091c5af4b09 100644 --- a/Mathlib/Analysis/Calculus/BumpFunction/FiniteDimension.lean +++ b/Mathlib/Analysis/Calculus/BumpFunction/FiniteDimension.lean @@ -93,7 +93,8 @@ theorem IsOpen.exists_smooth_support_eq {s : Set E} (hs : IsOpen s) : rcases exists_smooth_tsupport_subset (hs.mem_nhds hx) with ⟨f, hf⟩ let g : ι := ⟨f, (subset_tsupport f).trans hf.1, hf.2.1, hf.2.2.1, hf.2.2.2.1⟩ have : x ∈ support (g : E → ℝ) := by - simp only [hf.2.2.2.2, Subtype.coe_mk, mem_support, Ne, one_ne_zero, not_false_iff] + simp only [g, hf.2.2.2.2, Subtype.coe_mk, mem_support, Ne, one_ne_zero, + not_false_iff] exact mem_iUnion_of_mem _ this simp_rw [← this] apply isOpen_iUnion_countable diff --git a/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean b/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean index fe52bff69b40d..008902d921c30 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean @@ -152,12 +152,12 @@ theorem ContinuousLinearMap.norm_iteratedFDerivWithin_le_of_bilinear (B : E →L (ContinuousLinearMap.compL 𝕜 Fu G Gu (isoG.symm : G →L[𝕜] Gu)) Bu₀ := rfl have Bu_eq : (fun y => Bu (fu y) (gu y)) = isoG.symm ∘ (fun y => B (f y) (g y)) ∘ isoD := by ext1 y - simp [hBu, hBu₀, hfu, hgu] + simp [Du, Eu, Fu, Gu, hBu, hBu₀, hfu, hgu] -- All norms are preserved by the lifting process. have Bu_le : ‖Bu‖ ≤ ‖B‖ := by refine ContinuousLinearMap.opNorm_le_bound _ (norm_nonneg B) fun y => ?_ refine ContinuousLinearMap.opNorm_le_bound _ (by positivity) fun x => ?_ - simp only [hBu, hBu₀, compL_apply, coe_comp', Function.comp_apply, + simp only [Du, Eu, Fu, Gu, hBu, hBu₀, compL_apply, coe_comp', Function.comp_apply, ContinuousLinearEquiv.coe_coe, LinearIsometryEquiv.coe_coe, flip_apply, LinearIsometryEquiv.norm_map] calc diff --git a/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean b/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean index 06469937e9249..f9ce84b67909d 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean @@ -952,7 +952,7 @@ theorem HasFTaylorSeriesUpToOn.comp {n : WithTop ℕ∞} {g : F → G} {f : E hf.hasFDerivWithinAt (le_trans (mod_cast Nat.le_add_left 1 m) (ENat.add_one_natCast_le_withTop_of_lt hm)) hx convert HasFDerivWithinAt.linear_multilinear_comp (J.comp x K h) I B - simp only [Nat.succ_eq_add_one, Fintype.sum_option, comp_apply, faaDiBruno_aux1, + simp only [B, Nat.succ_eq_add_one, Fintype.sum_option, comp_apply, faaDiBruno_aux1, faaDiBruno_aux2] have B : HasFDerivWithinAt (fun x ↦ (q (f x)).taylorComp (p x) m) (∑ c : OrderedFinpartition m, ∑ i : Option (Fin c.length), diff --git a/Mathlib/Analysis/Calculus/FDeriv/Bilinear.lean b/Mathlib/Analysis/Calculus/FDeriv/Bilinear.lean index c4db82a10c282..67c304e3952ff 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Bilinear.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Bilinear.lean @@ -61,7 +61,7 @@ theorem IsBoundedBilinearMap.hasStrictFDerivAt (h : IsBoundedBilinearMap 𝕜 b) refine (isBigO_refl _ _).mul_isLittleO ((isLittleO_one_iff _).2 ?_) -- TODO: `continuity` fails exact (continuous_snd.fst.prod_mk continuous_fst.snd).norm.tendsto' _ _ (by simp) - _ = _ := by simp [Function.comp_def] + _ = _ := by simp [T, Function.comp_def] @[fun_prop] theorem IsBoundedBilinearMap.hasFDerivAt (h : IsBoundedBilinearMap 𝕜 b) (p : E × F) : diff --git a/Mathlib/Analysis/Calculus/LineDeriv/IntegrationByParts.lean b/Mathlib/Analysis/Calculus/LineDeriv/IntegrationByParts.lean index 07e351b377405..9a682b2461766 100644 --- a/Mathlib/Analysis/Calculus/LineDeriv/IntegrationByParts.lean +++ b/Mathlib/Analysis/Calculus/LineDeriv/IntegrationByParts.lean @@ -129,14 +129,14 @@ theorem integral_bilinear_hasLineDerivAt_right_eq_neg_left_of_integrable suffices H : ∫ (x : E' × ℝ), (B (f (L.symm x))) (g' (L.symm x)) ∂ν = -∫ (x : E' × ℝ), (B (f' (L.symm x))) (g (L.symm x)) ∂ν by have : μ = Measure.map L.symm ν := by - simp [Measure.map_map L.symm.continuous.measurable L.continuous.measurable] + simp [ν, Measure.map_map L.symm.continuous.measurable L.continuous.measurable] have hL : IsClosedEmbedding L.symm := L.symm.toHomeomorph.isClosedEmbedding simpa [this, hL.integral_map] using H have L_emb : MeasurableEmbedding L := L.toHomeomorph.measurableEmbedding apply integral_bilinear_hasLineDerivAt_right_eq_neg_left_of_integrable_aux2 - · simpa [L_emb.integrable_map_iff, Function.comp_def] using hf'g - · simpa [L_emb.integrable_map_iff, Function.comp_def] using hfg' - · simpa [L_emb.integrable_map_iff, Function.comp_def] using hfg + · simpa [ν, L_emb.integrable_map_iff, Function.comp_def] using hf'g + · simpa [ν, L_emb.integrable_map_iff, Function.comp_def] using hfg' + · simpa [ν, L_emb.integrable_map_iff, Function.comp_def] using hfg · intro x have : f = (f ∘ L.symm) ∘ (L : E →ₗ[ℝ] (E' × ℝ)) := by ext y; simp specialize hf (L.symm x) diff --git a/Mathlib/Analysis/Calculus/MeanValue.lean b/Mathlib/Analysis/Calculus/MeanValue.lean index 691da7a3e6562..0fcfbf75ac5d5 100644 --- a/Mathlib/Analysis/Calculus/MeanValue.lean +++ b/Mathlib/Analysis/Calculus/MeanValue.lean @@ -552,7 +552,7 @@ theorem norm_image_sub_le_of_norm_hasFDerivWithin_le' calc ‖f y - f x - φ (y - x)‖ = ‖f y - f x - (φ y - φ x)‖ := by simp _ = ‖f y - φ y - (f x - φ x)‖ := by congr 1; abel - _ = ‖g y - g x‖ := by simp + _ = ‖g y - g x‖ := by simp [g] _ ≤ C * ‖y - x‖ := Convex.norm_image_sub_le_of_norm_hasFDerivWithin_le hg bound hs xs ys /-- Variant of the mean value inequality on a convex set. Version with `fderivWithin`. -/ diff --git a/Mathlib/Analysis/Complex/Isometry.lean b/Mathlib/Analysis/Complex/Isometry.lean index 57ac151acc4e2..c2cf8a9bc727c 100644 --- a/Mathlib/Analysis/Complex/Isometry.lean +++ b/Mathlib/Analysis/Complex/Isometry.lean @@ -130,7 +130,7 @@ theorem linear_isometry_complex (f : ℂ ≃ₗᵢ[ℝ] ℂ) : ∃ a : Circle, f = rotation a ∨ f = conjLIE.trans (rotation a) := by let a : Circle := ⟨f 1, by simp [Submonoid.unitSphere, ← Complex.norm_eq_abs, f.norm_map]⟩ use a - have : (f.trans (rotation a).symm) 1 = 1 := by simpa using rotation_apply a⁻¹ (f 1) + have : (f.trans (rotation a).symm) 1 = 1 := by simpa [a] using rotation_apply a⁻¹ (f 1) refine (linear_isometry_complex_aux this).imp (fun h₁ => ?_) fun h₂ => ?_ · simpa using eq_mul_of_inv_mul_eq h₁ · exact eq_mul_of_inv_mul_eq h₂ diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean index 2925cad492d23..e152fc2ea4eac 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean @@ -110,7 +110,7 @@ theorem ModularGroup_T_zpow_mem_verticalStrip (z : ℍ) {N : ℕ} (hn : 0 < N) : refine ⟨?_, (by simp only [mul_neg, Int.cast_neg, Int.cast_mul, Int.cast_natCast, vadd_im, le_refl])⟩ have h : (N * (-n : ℝ) +ᵥ z).re = -N * Int.floor (z.re / N) + z.re := by - simp only [Int.cast_natCast, mul_neg, vadd_re, neg_mul] + simp only [n, Int.cast_natCast, mul_neg, vadd_re, neg_mul] norm_cast at * rw [h, add_comm] simp only [neg_mul, Int.cast_neg, Int.cast_mul, Int.cast_natCast] diff --git a/Mathlib/Analysis/Convex/Birkhoff.lean b/Mathlib/Analysis/Convex/Birkhoff.lean index 702bc496c5a0d..d69061ff3199b 100644 --- a/Mathlib/Analysis/Convex/Birkhoff.lean +++ b/Mathlib/Analysis/Convex/Birkhoff.lean @@ -48,7 +48,7 @@ private lemma exists_perm_eq_zero_implies_eq_zero [Nonempty n] {s : R} (hs : 0 < rw [exists_mem_doublyStochastic_eq_smul_iff hs.le] at hM let f (i : n) : Finset n := {j | M i j ≠ 0} have hf (A : Finset n) : #A ≤ #(A.biUnion f) := by - have (i) : ∑ j ∈ f i, M i j = s := by simp [sum_subset (filter_subset _ _), hM.2.1] + have (i) : ∑ j ∈ f i, M i j = s := by simp [f, sum_subset (filter_subset _ _), hM.2.1] have h₁ : ∑ i ∈ A, ∑ j ∈ f i, M i j = #A * s := by simp [this] have h₂ : ∑ i, ∑ j ∈ A.biUnion f, M i j = #(A.biUnion f) * s := by simp [sum_comm (t := A.biUnion f), hM.2.2, mul_comm s] @@ -110,7 +110,7 @@ private lemma doublyStochastic_sum_perm_aux (M : Matrix n n R) simp only [sub_apply, smul_apply, PEquiv.toMatrix_apply, Equiv.toPEquiv_apply, Option.mem_def, Option.some.injEq, smul_eq_mul, mul_ite, mul_one, mul_zero, sub_nonneg, sum_sub_distrib, sum_ite_eq, mem_univ, ↓reduceIte, N] - refine ⟨fun i' j => ?_, by simp [hM.2.1], by simp [← σ.eq_symm_apply, hM]⟩ + refine ⟨fun i' j => ?_, by simp [s', hM.2.1], by simp [s', ← σ.eq_symm_apply, hM]⟩ split case isTrue h => exact (hi' i' (by simp)).trans_eq (by rw [h]) case isFalse h => exact hM.1 _ _ diff --git a/Mathlib/Analysis/Convex/Combination.lean b/Mathlib/Analysis/Convex/Combination.lean index 777e221c7f434..983f2d7608437 100644 --- a/Mathlib/Analysis/Convex/Combination.lean +++ b/Mathlib/Analysis/Convex/Combination.lean @@ -285,7 +285,7 @@ theorem convexHull_range_eq_exists_affineCombination (v : ι → E) : convexHull · rintro i - by_cases hi : i ∈ s <;> by_cases hi' : i ∈ s' <;> simp [W, hi, hi', add_nonneg, mul_nonneg ha (hw₀ i _), mul_nonneg hb (hw₀' i _)] - · simp_rw [affineCombination_eq_linear_combination (s ∪ s') v _ hW₁, + · simp_rw [W, affineCombination_eq_linear_combination (s ∪ s') v _ hW₁, affineCombination_eq_linear_combination s v w hw₁, affineCombination_eq_linear_combination s' v w' hw₁', add_smul, sum_add_distrib] rw [← sum_subset subset_union_left, ← sum_subset subset_union_right] diff --git a/Mathlib/Analysis/Convex/Cone/Extension.lean b/Mathlib/Analysis/Convex/Cone/Extension.lean index 31d7448ecb98e..31471bbcab337 100644 --- a/Mathlib/Analysis/Convex/Cone/Extension.lean +++ b/Mathlib/Analysis/Convex/Cone/Extension.lean @@ -68,7 +68,7 @@ theorem step (nonneg : ∀ x : f.domain, (x : E) ∈ s → 0 ≤ f x) set Sp := f '' { x : f.domain | (x : E) + y ∈ s } set Sn := f '' { x : f.domain | -(x : E) - y ∈ s } suffices (upperBounds Sn ∩ lowerBounds Sp).Nonempty by - simpa only [Set.Nonempty, upperBounds, lowerBounds, forall_mem_image] using this + simpa only [Sp, Sn, Set.Nonempty, upperBounds, lowerBounds, forall_mem_image] using this refine exists_between_of_forall_le (Nonempty.image f ?_) (Nonempty.image f (dense y)) ?_ · rcases dense (-y) with ⟨x, hx⟩ rw [← neg_neg x, NegMemClass.coe_neg, ← sub_eq_add_neg] at hx diff --git a/Mathlib/Analysis/Convex/Jensen.lean b/Mathlib/Analysis/Convex/Jensen.lean index 044b54d8b3d24..332967af25195 100644 --- a/Mathlib/Analysis/Convex/Jensen.lean +++ b/Mathlib/Analysis/Convex/Jensen.lean @@ -110,7 +110,7 @@ lemma StrictConvexOn.map_sum_lt (hf : StrictConvexOn 𝕜 s f) (h₀ : ∀ i ∈ have hk : k ∉ u := by simp [u] have ht : t = (u.cons k hk).cons j (mem_cons.not.2 <| not_or_intro (ne_of_apply_ne _ hjk) hj) := by - simp [insert_erase this, insert_erase ‹j ∈ t›, *] + simp [u, insert_erase this, insert_erase ‹j ∈ t›, *] clear_value u subst ht simp only [sum_cons] diff --git a/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean b/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean index 08b29d45a4d25..367aebfa1d6f8 100644 --- a/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean +++ b/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean @@ -116,7 +116,7 @@ theorem tendsto_integral_exp_inner_smul_cocompact_of_continuous_compact_support suffices A = Metric.closedBall (0 : V) (R + 1) by rw [this] exact Metric.isClosed_ball.measurableSet - simp_rw [Metric.closedBall, dist_eq_norm, sub_zero] + simp_rw [A, Metric.closedBall, dist_eq_norm, sub_zero] obtain ⟨B, hB_pos, hB_vol⟩ : ∃ B : ℝ≥0, 0 < B ∧ volume A ≤ B := by have hc : IsCompact A := by simpa only [Metric.closedBall, dist_eq_norm, sub_zero] using isCompact_closedBall (0 : V) _ diff --git a/Mathlib/Analysis/InnerProductSpace/PiL2.lean b/Mathlib/Analysis/InnerProductSpace/PiL2.lean index 3469e47591c82..4813828390c93 100644 --- a/Mathlib/Analysis/InnerProductSpace/PiL2.lean +++ b/Mathlib/Analysis/InnerProductSpace/PiL2.lean @@ -926,7 +926,8 @@ noncomputable def LinearIsometry.extend (L : S →ₗᵢ[𝕜] V) : V →ₗᵢ[ calc finrank 𝕜 LSᗮ = finrank 𝕜 V - finrank 𝕜 LS := by simp only [← LS.finrank_add_finrank_orthogonal, add_tsub_cancel_left] - _ = finrank 𝕜 V - finrank 𝕜 S := by simp only [LinearMap.finrank_range_of_inj L.injective] + _ = finrank 𝕜 V - finrank 𝕜 S := by + simp only [LS, LinearMap.finrank_range_of_inj L.injective] _ = finrank 𝕜 Sᗮ := by simp only [← S.finrank_add_finrank_orthogonal, add_tsub_cancel_left] exact @@ -952,8 +953,8 @@ noncomputable def LinearIsometry.extend (L : S →ₗᵢ[𝕜] V) : V →ₗᵢ[ have Lp1x : L (p1 x) ∈ LinearMap.range L.toLinearMap := LinearMap.mem_range_self L.toLinearMap (p1 x) have Lp2x : L3 (p2 x) ∈ (LinearMap.range L.toLinearMap)ᗮ := by - simp only [LinearIsometry.coe_comp, Function.comp_apply, Submodule.coe_subtypeₗᵢ, ← - Submodule.range_subtype LSᗮ] + simp only [LinearIsometry.coe_comp, Function.comp_apply, Submodule.coe_subtypeₗᵢ, + ← Submodule.range_subtype LSᗮ] apply LinearMap.mem_range_self apply Submodule.inner_right_of_mem_orthogonal Lp1x Lp2x -- Apply the Pythagorean theorem and simplify diff --git a/Mathlib/Analysis/InnerProductSpace/Projection.lean b/Mathlib/Analysis/InnerProductSpace/Projection.lean index 6be13bb9ce5f1..902821ff1cdfd 100644 --- a/Mathlib/Analysis/InnerProductSpace/Projection.lean +++ b/Mathlib/Analysis/InnerProductSpace/Projection.lean @@ -1314,7 +1314,7 @@ theorem maximal_orthonormal_iff_orthogonalComplement_eq_bot (hv : Orthonormal rintro ⟨x, hx', hx⟩ -- take a nonzero vector and normalize it let e := (‖x‖⁻¹ : 𝕜) • x - have he : ‖e‖ = 1 := by simp [norm_smul_inv_norm hx] + have he : ‖e‖ = 1 := by simp [e, norm_smul_inv_norm hx] have he' : e ∈ (span 𝕜 v)ᗮ := smul_mem' _ _ hx' have he'' : e ∉ v := by intro hev diff --git a/Mathlib/Analysis/Normed/Group/ControlledClosure.lean b/Mathlib/Analysis/Normed/Group/ControlledClosure.lean index 2a79581ba1dcc..62167b8250217 100644 --- a/Mathlib/Analysis/Normed/Group/ControlledClosure.lean +++ b/Mathlib/Analysis/Normed/Group/ControlledClosure.lean @@ -59,7 +59,7 @@ theorem controlled_closure_of_complete {f : NormedAddGroupHom G H} {K : AddSubgr calc ‖u n‖ ≤ C * ‖v n‖ := hnorm_u n _ ≤ C * b n := by gcongr; exact (hv _ <| Nat.succ_le_iff.mp hn).le - _ = (1 / 2) ^ n * (ε * ‖h‖ / 2) := by simp [mul_div_cancel₀ _ hC.ne.symm] + _ = (1 / 2) ^ n * (ε * ‖h‖ / 2) := by simp [b, mul_div_cancel₀ _ hC.ne.symm] _ = ε * ‖h‖ / 2 * (1 / 2) ^ n := mul_comm _ _ -- We now show that the limit `g` of `s` is the desired preimage. obtain ⟨g : G, hg⟩ := cauchySeq_tendsto_of_complete this @@ -89,7 +89,7 @@ theorem controlled_closure_of_complete {f : NormedAddGroupHom G H} {K : AddSubgr have : (∑ k ∈ range (n + 1), C * b k) ≤ ε * ‖h‖ := calc (∑ k ∈ range (n + 1), C * b k) _ = (∑ k ∈ range (n + 1), (1 / 2 : ℝ) ^ k) * (ε * ‖h‖ / 2) := by - simp only [mul_div_cancel₀ _ hC.ne.symm, ← sum_mul] + simp only [b, mul_div_cancel₀ _ hC.ne.symm, ← sum_mul] _ ≤ 2 * (ε * ‖h‖ / 2) := by gcongr; apply sum_geometric_two_le _ = ε * ‖h‖ := mul_div_cancel₀ _ two_ne_zero calc diff --git a/Mathlib/Analysis/Normed/Operator/Banach.lean b/Mathlib/Analysis/Normed/Operator/Banach.lean index 8810020324216..27ca8d21c6ddf 100644 --- a/Mathlib/Analysis/Normed/Operator/Banach.lean +++ b/Mathlib/Analysis/Normed/Operator/Banach.lean @@ -121,7 +121,7 @@ theorem exists_approx_preimage_norm_le (surj : Surjective f) : calc ‖f x - d • y‖ = ‖f x₁ - (a + d • y) - (f x₂ - a)‖ := by congr 1 - simp only [f.map_sub] + simp only [x, f.map_sub] abel _ ≤ ‖f x₁ - (a + d • y)‖ + ‖f x₂ - a‖ := norm_sub_le _ _ _ ≤ δ + δ := by rw [dist_eq_norm'] at h₁ h₂; gcongr diff --git a/Mathlib/Analysis/NormedSpace/RieszLemma.lean b/Mathlib/Analysis/NormedSpace/RieszLemma.lean index 553d89feaaa42..1fc30104864f2 100644 --- a/Mathlib/Analysis/NormedSpace/RieszLemma.lean +++ b/Mathlib/Analysis/NormedSpace/RieszLemma.lean @@ -98,7 +98,7 @@ theorem riesz_lemma_of_norm_lt {c : 𝕜} (hc : 1 < ‖c‖) {R : ℝ} (hR : ‖ _ = ‖d‖ * (‖c‖ / R * ‖x‖) := by simp only [norm_smul] ring - _ ≤ ‖d‖ * ‖x - y'‖ := by gcongr; exact hx y' (by simp [Submodule.smul_mem _ _ hy]) + _ ≤ ‖d‖ * ‖x - y'‖ := by gcongr; exact hx y' (by simp [y', Submodule.smul_mem _ _ hy]) _ = ‖d • x - y‖ := by rw [yy', ← smul_sub, norm_smul] theorem Metric.closedBall_infDist_compl_subset_closure {x : F} {s : Set F} (hx : x ∈ s) : diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/LogBounds.lean b/Mathlib/Analysis/SpecialFunctions/Complex/LogBounds.lean index 51294919f4f57..061caf910fee5 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/LogBounds.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/LogBounds.lean @@ -148,7 +148,8 @@ lemma norm_log_sub_logTaylor_le (n : ℕ) {z : ℂ} (hz : ‖z‖ < 1) : have H : f z = z * ∫ t in (0 : ℝ)..1, (-(t * z)) ^ n * (1 + t * z)⁻¹ := by convert (integral_unitInterval_deriv_eq_sub hcont hderiv).symm using 1 · simp only [f, zero_add, add_zero, log_one, logTaylor_at_zero, sub_self, sub_zero] - · simp only [add_zero, log_one, logTaylor_at_zero, sub_self, real_smul, zero_add, smul_eq_mul] + · simp only [f', add_zero, log_one, logTaylor_at_zero, sub_self, real_smul, zero_add, + smul_eq_mul] unfold f at H simp only [H, norm_mul] simp_rw [neg_pow (_ * z) n, mul_assoc, intervalIntegral.integral_const_mul, mul_pow, diff --git a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/PosPart.lean b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/PosPart.lean index 6c5d5082a67e1..b025d0d6c2621 100644 --- a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/PosPart.lean +++ b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/PosPart.lean @@ -236,7 +236,7 @@ lemma posPart_negPart_unique {a b c : A} (habc : a = b - c) (hbc : b * c = 0) all_goals exact isCompact_quasispectrum _ obtain ⟨has, hbs, hcs⟩ : σₙ ℝ a ⊆ s ∧ σₙ ℝ b ⊆ s ∧ σₙ ℝ (-c) ⊆ s := by refine ⟨?_, ?_, ?_⟩; all_goals intro; aesop - let _ : Zero s := ⟨0, by aesop⟩ + let zero : Zero s := ⟨0, by aesop⟩ have s0 : (0 : s) = (0 : ℝ) := rfl /- The continuous functional calculi for functions `f g : C(s, ℝ)₀` applied to `b` and `(-c)` are orthogonal (i.e., the product is always zero). -/ @@ -248,8 +248,8 @@ lemma posPart_negPart_unique {a b c : A} (habc : a = b - c) (hbc : b * c = 0) all_goals refine g.mul_nonUnitalStarAlgHom_apply_eq_zero s0 _ _ ?_ ?_ (cfcₙHomSuperset_continuous hc' hcs) - all_goals simp only [star_trivial, cfcₙHomSuperset_id' hb' hbs, cfcₙHomSuperset_id' hc' hcs, - mul_neg, hbc, neg_zero] + all_goals simp only [zero, star_trivial, cfcₙHomSuperset_id' hb' hbs, + cfcₙHomSuperset_id' hc' hcs, mul_neg, hbc, neg_zero] have mul₂ (f g : C(s, ℝ)₀) : (cfcₙHomSuperset hc' hcs f) * (cfcₙHomSuperset hb' hbs g) = 0 := by simpa only [star_mul, star_zero, ← map_star, star_trivial] using congr(star $(mul₁ g f)) /- `fun f ↦ cfcₙ f b + cfcₙ f (-c)` defines a star homomorphism `ψ : C(s, ℝ)₀ →⋆ₙₐ[ℝ] A` which @@ -261,13 +261,14 @@ lemma posPart_negPart_unique {a b c : A} (habc : a = b - c) (hbc : b * c = 0) toFun := cfcₙHomSuperset hb' hbs + cfcₙHomSuperset hc' hcs map_zero' := by simp [-cfcₙHomSuperset_apply] map_mul' := fun f g ↦ by - simp only [Pi.add_apply, map_mul, mul_add, add_mul, mul₂, add_zero, mul₁, zero_add] + simp only [zero, Pi.add_apply, map_mul, mul_add, add_mul, mul₂, add_zero, mul₁, + zero_add] map_star' := fun f ↦ by simp [← map_star] } have key : (cfcₙHomSuperset ha has) = ψ := UniqueNonUnitalContinuousFunctionalCalculus.eq_of_continuous_of_map_id s rfl (cfcₙHomSuperset ha has) ψ (cfcₙHomSuperset_continuous ha has) ((cfcₙHomSuperset_continuous hb' hbs).add (cfcₙHomSuperset_continuous hc' hcs)) - (by simpa [ψ, -cfcₙHomSuperset_apply, cfcₙHomSuperset_id, sub_eq_add_neg] using habc) + (by simpa [zero, ψ, -cfcₙHomSuperset_apply, cfcₙHomSuperset_id, sub_eq_add_neg] using habc) /- Applying the equality of star homomorphisms to the function `(·⁺ : ℝ → ℝ)` we find that `b = cfcₙ id b + cfcₙ 0 (-c) = cfcₙ (·⁺) b - cfcₙ (·⁺) (-c) = cfcₙ (·⁺) a = a⁺`, where the second equality follows because these functions are equal on the spectra of `b` and `-c`, @@ -284,8 +285,9 @@ lemma posPart_negPart_unique {a b c : A} (habc : a = b - c) (hbc : b * c = 0) all_goals refine cfcₙ_congr fun x hx ↦ Eq.symm ?_ lift x to σₙ ℝ _ using hx - simp only [Subtype.val_injective.extend_apply, comp_apply, coe_mk, ContinuousMap.coe_mk, - Subtype.map_coe, id_eq, _root_.posPart_eq_self, f, Pi.zero_apply, posPart_eq_zero] + simp only [zero, Subtype.val_injective.extend_apply, comp_apply, coe_mk, + ContinuousMap.coe_mk, Subtype.map_coe, id_eq, _root_.posPart_eq_self, f, Pi.zero_apply, + posPart_eq_zero] · exact quasispectrum_nonneg_of_nonneg b hb x.val x.property · obtain ⟨x, hx⟩ := x simp only [← neg_nonneg] @@ -297,7 +299,7 @@ lemma posPart_negPart_unique {a b c : A} (habc : a = b - c) (hbc : b * c = 0) _ = a⁺ := by refine cfcₙ_congr fun x hx ↦ ?_ lift x to σₙ ℝ a using hx - simp [Subtype.val_injective.extend_apply, f] + simp [zero, Subtype.val_injective.extend_apply, f] end CFC diff --git a/Mathlib/Analysis/SpecialFunctions/Integrals.lean b/Mathlib/Analysis/SpecialFunctions/Integrals.lean index ed9d8bc8c55e9..f80d1afa00a62 100644 --- a/Mathlib/Analysis/SpecialFunctions/Integrals.lean +++ b/Mathlib/Analysis/SpecialFunctions/Integrals.lean @@ -664,7 +664,7 @@ theorem integral_cos_pow_aux : · calc (∫ x in a..b, cos x ^ (n + 2)) = ∫ x in a..b, cos x ^ (n + 1) * cos x := by simp only [_root_.pow_succ] - _ = C + (n + 1) * ∫ x in a..b, sin x ^ 2 * cos x ^ n := by simp [H, h, sq, -neg_add_rev] + _ = C + (n + 1) * ∫ x in a..b, sin x ^ 2 * cos x ^ n := by simp [C, H, h, sq, -neg_add_rev] _ = C + (n + 1) * ∫ x in a..b, cos x ^ n - cos x ^ (n + 2) := by simp [sin_sq, sub_mul, ← pow_add, add_comm] _ = (C + (n + 1) * ∫ x in a..b, cos x ^ n) - (n + 1) * ∫ x in a..b, cos x ^ (n + 2) := by diff --git a/Mathlib/CategoryTheory/Abelian/EpiWithInjectiveKernel.lean b/Mathlib/CategoryTheory/Abelian/EpiWithInjectiveKernel.lean index 40e7feb1d78af..5c34bd58bd141 100644 --- a/Mathlib/CategoryTheory/Abelian/EpiWithInjectiveKernel.lean +++ b/Mathlib/CategoryTheory/Abelian/EpiWithInjectiveKernel.lean @@ -43,7 +43,7 @@ lemma epiWithInjectiveKernel_iff {X Y : C} (g : X ⟶ Y) : exact ⟨_, inferInstance, _, S.zero, ⟨ShortComplex.Splitting.ofExactOfRetraction S (S.exact_of_f_is_kernel (kernelIsKernel g)) (Injective.factorThru (𝟙 _) (kernel.ι g)) - (by simp) inferInstance⟩⟩ + (by simp [S]) inferInstance⟩⟩ · rintro ⟨I, _, f, w, ⟨σ⟩⟩ have : IsSplitEpi g := ⟨σ.s, σ.s_g⟩ let e : I ≅ kernel g := diff --git a/Mathlib/CategoryTheory/Extensive.lean b/Mathlib/CategoryTheory/Extensive.lean index 87488121b4c56..dd21c2f557352 100644 --- a/Mathlib/CategoryTheory/Extensive.lean +++ b/Mathlib/CategoryTheory/Extensive.lean @@ -102,7 +102,7 @@ theorem FinitaryExtensive.vanKampen [FinitaryExtensive C] {F : Discrete WalkingP have : F = pair X Y := by apply Functor.hext · rintro ⟨⟨⟩⟩ <;> rfl - · rintro ⟨⟨⟩⟩ ⟨j⟩ ⟨⟨rfl : _ = j⟩⟩ <;> simp + · rintro ⟨⟨⟩⟩ ⟨j⟩ ⟨⟨rfl : _ = j⟩⟩ <;> simp [X, Y] clear_value X Y subst this exact FinitaryExtensive.van_kampen' c hc @@ -491,7 +491,7 @@ lemma FinitaryPreExtensive.hasPullbacks_of_is_coproduct [FinitaryPreExtensive C] let e' : c.pt ≅ f i ⨿ (∐ fun j : ({i}ᶜ : Set ι) ↦ f j) := hc.coconePointUniqueUpToIso (getColimitCocone _).2 ≪≫ e have : coprod.inl ≫ e'.inv = c.ι.app ⟨i⟩ := by - simp only [e', Iso.trans_inv, coprod.desc_comp, colimit.ι_desc, BinaryCofan.mk_pt, + simp only [e, e', Iso.trans_inv, coprod.desc_comp, colimit.ι_desc, BinaryCofan.mk_pt, BinaryCofan.ι_app_left, BinaryCofan.mk_inl] exact colimit.comp_coconePointUniqueUpToIso_inv _ _ clear_value e' diff --git a/Mathlib/CategoryTheory/GradedObject/Trifunctor.lean b/Mathlib/CategoryTheory/GradedObject/Trifunctor.lean index 3950b5d3d1b1a..d1c4c259169b5 100644 --- a/Mathlib/CategoryTheory/GradedObject/Trifunctor.lean +++ b/Mathlib/CategoryTheory/GradedObject/Trifunctor.lean @@ -345,7 +345,7 @@ noncomputable def isColimitCofan₃MapBifunctor₁₂BifunctorMapObj (j : J) : refine IsColimit.ofIsoColimit (isColimitCofanMapObjComp Z p' ρ₁₂.q r ρ₁₂.hpq j (fun ⟨i₁₂, i₃⟩ h => c₁₂'' ⟨⟨i₁₂, i₃⟩, h⟩) (fun ⟨i₁₂, i₃⟩ h => h₁₂'' ⟨⟨i₁₂, i₃⟩, h⟩) c hc) (Cocones.ext (Iso.refl _) (fun ⟨⟨i₁, i₂, i₃⟩, h⟩ => ?_)) - dsimp [Cofan.inj, c₁₂'', Z] + dsimp [Cofan.inj, c₁₂'', Z, p'] rw [comp_id, Functor.map_id, id_comp] rfl @@ -523,7 +523,7 @@ noncomputable def isColimitCofan₃MapBifunctorBifunctor₂₃MapObj (j : J) : refine IsColimit.ofIsoColimit (isColimitCofanMapObjComp Z p' ρ₂₃.q r ρ₂₃.hpq j (fun ⟨i₁, i₂₃⟩ h => c₂₃'' ⟨⟨i₁, i₂₃⟩, h⟩) (fun ⟨i₁, i₂₃⟩ h => h₂₃'' ⟨⟨i₁, i₂₃⟩, h⟩) c hc) (Cocones.ext (Iso.refl _) (fun ⟨⟨i₁, i₂, i₃⟩, h⟩ => ?_)) - dsimp [Cofan.inj, c₂₃''] + dsimp [Cofan.inj, c₂₃'', Z, p', e] rw [comp_id, id_comp] rfl diff --git a/Mathlib/CategoryTheory/Idempotents/FunctorCategories.lean b/Mathlib/CategoryTheory/Idempotents/FunctorCategories.lean index aa1af4a9269f3..29b94beac78e3 100644 --- a/Mathlib/CategoryTheory/Idempotents/FunctorCategories.lean +++ b/Mathlib/CategoryTheory/Idempotents/FunctorCategories.lean @@ -70,14 +70,14 @@ instance functor_category_isIdempotentComplete [IsIdempotentComplete C] : let e : F ⟶ Y := { app := fun j => equalizer.lift (p.app j) (by simpa only [comp_id] using (congr_app hp j).symm) - naturality := fun j j' φ => equalizer.hom_ext (by simp) } + naturality := fun j j' φ => equalizer.hom_ext (by simp [Y]) } use Y, i, e constructor · ext j dsimp rw [assoc, equalizer.lift_ι, ← equalizer.condition, id_comp, comp_id] · ext j - simp + simp [Y, i, e] namespace KaroubiFunctorCategoryEmbedding variable {J C} diff --git a/Mathlib/CategoryTheory/Limits/Constructions/LimitsOfProductsAndEqualizers.lean b/Mathlib/CategoryTheory/Limits/Constructions/LimitsOfProductsAndEqualizers.lean index 353032099523c..84ff0c9c14b61 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/LimitsOfProductsAndEqualizers.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/LimitsOfProductsAndEqualizers.lean @@ -166,19 +166,19 @@ lemma preservesLimit_of_preservesEqualizers_and_product : let I := equalizer s t let i : I ⟶ P := equalizer.ι s t apply preservesLimit_of_preserves_limit_cone - (buildIsLimit s t (by simp [s]) (by simp [t]) (limit.isLimit _) (limit.isLimit _) - (limit.isLimit _)) + (buildIsLimit s t (by simp [P, s]) (by simp [P, t]) (limit.isLimit _) + (limit.isLimit _) (limit.isLimit _)) apply IsLimit.ofIsoLimit (buildIsLimit _ _ _ _ _ _ _) _ · exact Fan.mk _ fun j => G.map (Pi.π _ j) · exact Fan.mk (G.obj Q) fun f => G.map (Pi.π _ f) · apply G.map s · apply G.map t · intro f - dsimp [s, Fan.mk] + dsimp [P, Q, s, Fan.mk] simp only [← G.map_comp, limit.lift_π] congr · intro f - dsimp [t, Fan.mk] + dsimp [P, Q, t, Fan.mk] simp only [← G.map_comp, limit.lift_π] apply congrArg G.map dsimp @@ -191,7 +191,7 @@ lemma preservesLimit_of_preservesEqualizers_and_product : · apply isLimitForkMapOfIsLimit apply equalizerIsEqualizer · refine Cones.ext (Iso.refl _) ?_ - intro j; dsimp; simp + intro j; dsimp [P, Q, I, i]; simp -- See note [dsimp, simp]. end @@ -382,8 +382,8 @@ lemma preservesColimit_of_preservesCoequalizers_and_coproduct : let I := coequalizer s t let i : P ⟶ I := coequalizer.π s t apply preservesColimit_of_preserves_colimit_cocone - (buildIsColimit s t (by simp [s]) (by simp [t]) (colimit.isColimit _) (colimit.isColimit _) - (colimit.isColimit _)) + (buildIsColimit s t (by simp [P, s]) (by simp [P, t]) (colimit.isColimit _) + (colimit.isColimit _) (colimit.isColimit _)) apply IsColimit.ofIsoColimit (buildIsColimit _ _ _ _ _ _ _) _ · refine Cofan.mk (G.obj Q) fun j => G.map ?_ apply Sigma.ι _ j @@ -392,11 +392,11 @@ lemma preservesColimit_of_preservesCoequalizers_and_coproduct : · apply G.map s · apply G.map t · intro f - dsimp [s, Cofan.mk] + dsimp [P, Q, s, Cofan.mk] simp only [← G.map_comp, colimit.ι_desc] congr · intro f - dsimp [t, Cofan.mk] + dsimp [P, Q, t, Cofan.mk] simp only [← G.map_comp, colimit.ι_desc] dsimp · refine Cofork.ofπ (G.map i) ?_ @@ -409,7 +409,7 @@ lemma preservesColimit_of_preservesCoequalizers_and_coproduct : apply coequalizerIsCoequalizer refine Cocones.ext (Iso.refl _) ?_ intro j - dsimp + dsimp [P, Q, I, i] simp -- See note [dsimp, simp]. diff --git a/Mathlib/CategoryTheory/Limits/Constructions/Pullbacks.lean b/Mathlib/CategoryTheory/Limits/Constructions/Pullbacks.lean index e6c2252ce3d8d..577fbad322d22 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/Pullbacks.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/Pullbacks.lean @@ -31,7 +31,9 @@ theorem hasLimit_cospan_of_hasLimit_pair_of_hasLimit_parallelPair {C : Type u} [ let e := equalizer.ι (π₁ ≫ f) (π₂ ≫ g) HasLimit.mk { cone := - PullbackCone.mk (e ≫ π₁) (e ≫ π₂) <| by rw [Category.assoc, equalizer.condition]; simp + PullbackCone.mk (e ≫ π₁) (e ≫ π₂) <| by + rw [Category.assoc, equalizer.condition] + simp [e] isLimit := PullbackCone.IsLimit.mk _ (fun s => equalizer.lift (prod.lift (s.π.app WalkingCospan.left) (s.π.app WalkingCospan.right)) <| by diff --git a/Mathlib/CategoryTheory/Limits/FilteredColimitCommutesProduct.lean b/Mathlib/CategoryTheory/Limits/FilteredColimitCommutesProduct.lean index ac195d5acac16..790a96b3416e6 100644 --- a/Mathlib/CategoryTheory/Limits/FilteredColimitCommutesProduct.lean +++ b/Mathlib/CategoryTheory/Limits/FilteredColimitCommutesProduct.lean @@ -136,9 +136,9 @@ theorem Types.isIso_colimitPointwiseProductToProductColimit (F : ∀ i, I i ⥤ let yk' : (pointwiseProduct F).obj k := (pointwiseProduct F).map (IsFiltered.rightToMax ky ky') yk₀' obtain rfl : y = colimit.ι (pointwiseProduct F) k yk := by - simp only [yk, Types.Colimit.w_apply', hyk₀] + simp only [k, yk, Types.Colimit.w_apply', hyk₀] obtain rfl : y' = colimit.ι (pointwiseProduct F) k yk' := by - simp only [yk', Types.Colimit.w_apply', hyk₀'] + simp only [k, yk', Types.Colimit.w_apply', hyk₀'] dsimp only [pointwiseProduct_obj] at yk yk' have hch : ∀ (s : α), ∃ (i' : I s) (hi' : k s ⟶ i'), (F s).map hi' (Pi.π (fun s => (F s).obj (k s)) s yk) = diff --git a/Mathlib/CategoryTheory/Limits/SmallComplete.lean b/Mathlib/CategoryTheory/Limits/SmallComplete.lean index 66ba7ba7e12d4..41ff160e8b730 100644 --- a/Mathlib/CategoryTheory/Limits/SmallComplete.lean +++ b/Mathlib/CategoryTheory/Limits/SmallComplete.lean @@ -61,10 +61,10 @@ instance (priority := 100) : Quiver.IsThin C := fun X Y => refine ⟨⟨Pi.lift, fun f k => f ≫ Pi.π _ k, ?_, ?_⟩⟩ · intro f ext k - simp + simp [yp] · intro f ext ⟨j⟩ - simp + simp [yp] · apply Cardinal.mk_le_of_injective _ · intro f exact ⟨_, _, f⟩ diff --git a/Mathlib/CategoryTheory/Limits/VanKampen.lean b/Mathlib/CategoryTheory/Limits/VanKampen.lean index 75e50cc0928fc..705c13f505a00 100644 --- a/Mathlib/CategoryTheory/Limits/VanKampen.lean +++ b/Mathlib/CategoryTheory/Limits/VanKampen.lean @@ -352,7 +352,7 @@ theorem IsUniversalColimit.map_reflective · intro j rw [← Category.assoc, Iso.comp_inv_eq] ext - all_goals simp only [PreservesPullback.iso_hom_fst, PreservesPullback.iso_hom_snd, + all_goals simp only [c'', PreservesPullback.iso_hom_fst, PreservesPullback.iso_hom_snd, pullback.lift_fst, pullback.lift_snd, Category.assoc, Functor.mapCocone_ι_app, ← Gl.map_comp] · rw [IsIso.comp_inv_eq, adj.counit_naturality] @@ -363,7 +363,8 @@ theorem IsUniversalColimit.map_reflective rw [Category.comp_id, Category.assoc] have : cf.hom ≫ (PreservesPullback.iso _ _ _).hom ≫ pullback.fst _ _ ≫ adj.counit.app _ = 𝟙 _ := by - simp only [IsIso.inv_hom_id, Iso.inv_hom_id_assoc, Category.assoc, pullback.lift_fst_assoc] + simp only [cf, IsIso.inv_hom_id, Iso.inv_hom_id_assoc, Category.assoc, + pullback.lift_fst_assoc] have : IsIso cf := by apply @Cocones.cocone_iso_of_hom_iso (i := ?_) rw [← IsIso.eq_comp_inv] at this @@ -373,12 +374,12 @@ theorem IsUniversalColimit.map_reflective · exact ⟨IsColimit.precomposeHomEquiv β c' <| (isColimitOfPreserves Gl Hc'').ofIsoColimit (asIso cf).symm⟩ · ext j - dsimp + dsimp [c''] simp only [Category.comp_id, Category.id_comp, Category.assoc, Functor.map_comp, pullback.lift_snd] · intro j apply IsPullback.of_right _ _ (IsPullback.of_hasPullback _ _) - · dsimp [α'] + · dsimp [α', c''] simp only [Category.comp_id, Category.id_comp, Category.assoc, Functor.map_comp, pullback.lift_fst] rw [← Category.comp_id (Gr.map f)] @@ -390,7 +391,7 @@ theorem IsUniversalColimit.map_reflective dsimp simp only [Category.comp_id, Adjunction.right_triangle_components, Category.id_comp, Category.assoc] - · dsimp + · dsimp [c''] simp only [Category.comp_id, Category.id_comp, Category.assoc, Functor.map_comp, pullback.lift_snd] @@ -506,7 +507,7 @@ theorem BinaryCofan.isVanKampen_iff (c : BinaryCofan X Y) : have : F' = pair X' Y' := by apply Functor.hext · rintro ⟨⟨⟩⟩ <;> rfl - · rintro ⟨⟨⟩⟩ ⟨j⟩ ⟨⟨rfl : _ = j⟩⟩ <;> simp + · rintro ⟨⟨⟩⟩ ⟨j⟩ ⟨⟨rfl : _ = j⟩⟩ <;> simp [X', Y'] clear_value X' Y' subst this change BinaryCofan X' Y' at c' diff --git a/Mathlib/CategoryTheory/Localization/CalculusOfFractions.lean b/Mathlib/CategoryTheory/Localization/CalculusOfFractions.lean index ab6a1a44daa7d..a27c2821017c6 100644 --- a/Mathlib/CategoryTheory/Localization/CalculusOfFractions.lean +++ b/Mathlib/CategoryTheory/Localization/CalculusOfFractions.lean @@ -385,9 +385,9 @@ noncomputable def Hom.comp {X Y Z : C} (z₁ : Hom W X Y) (z₂ : Hom W Y Z) : H obtain ⟨Z, u, hu, fac₃⟩ := HasLeftCalculusOfFractions.ext _ _ _ a₁.hs eq simp only [assoc] at fac₃ refine ⟨Z, w₁.s ≫ u, u, ?_, ?_, ?_⟩ - · dsimp + · dsimp [p₁] simp only [assoc] - · dsimp + · dsimp [p₁] simp only [assoc, fac₃] · dsimp simp only [assoc] @@ -401,9 +401,9 @@ noncomputable def Hom.comp {X Y Z : C} (z₁ : Hom W X Y) (z₂ : Hom W Y Z) : H obtain ⟨Z, u, hu, fac₄⟩ := HasLeftCalculusOfFractions.ext _ _ _ a₁.hs eq simp only [assoc] at fac₄ refine ⟨Z, q.f ≫ u, q.s ≫ u, ?_, ?_, ?_⟩ - · simp only [assoc, reassoc_of% fac₃] + · simp only [p₁, p₂, assoc, reassoc_of% fac₃] · rw [assoc, assoc, assoc, assoc, fac₄, reassoc_of% hft] - · simp only [assoc, ← reassoc_of% fac₃] + · simp only [p₁, p₂, assoc, ← reassoc_of% fac₃] exact W.comp_mem _ _ b.hs (W.comp_mem _ _ z₂.hs (W.comp_mem _ _ w₂.hs (W.comp_mem _ _ q.hs hu))) · have eq : a₂.s ≫ z₂.f ≫ w₂.s = a₂.s ≫ t₂ ≫ w₂.f := by @@ -411,11 +411,11 @@ noncomputable def Hom.comp {X Y Z : C} (z₁ : Hom W X Y) (z₂ : Hom W Y Z) : H obtain ⟨Z, u, hu, fac₄⟩ := HasLeftCalculusOfFractions.ext _ _ _ a₂.hs eq simp only [assoc] at fac₄ refine ⟨Z, u, w₂.s ≫ u, ?_, ?_, ?_⟩ - · dsimp + · dsimp [p₁, p₂] simp only [assoc] - · dsimp + · dsimp [p₁, p₂] simp only [assoc, fac₄] - · dsimp + · dsimp [p₁, p₂] simp only [assoc] exact W.comp_mem _ _ b.hs (W.comp_mem _ _ z₂.hs (W.comp_mem _ _ w₂.hs hu)) diff --git a/Mathlib/CategoryTheory/Monoidal/Functor.lean b/Mathlib/CategoryTheory/Monoidal/Functor.lean index a2388f08a6793..83192bd9443b1 100644 --- a/Mathlib/CategoryTheory/Monoidal/Functor.lean +++ b/Mathlib/CategoryTheory/Monoidal/Functor.lean @@ -972,7 +972,7 @@ instance [e.functor.Monoidal] : e.symm.inverse.Monoidal := inferInstanceAs (e.fu noncomputable def inverseMonoidal [e.functor.Monoidal] : e.inverse.Monoidal := by letI := e.toAdjunction.rightAdjointLaxMonoidal have : IsIso (LaxMonoidal.ε e.inverse) := by - simp only [Adjunction.rightAdjointLaxMonoidal_ε, Adjunction.homEquiv_unit] + simp only [this, Adjunction.rightAdjointLaxMonoidal_ε, Adjunction.homEquiv_unit] infer_instance have : ∀ (X Y : D), IsIso (LaxMonoidal.μ e.inverse X Y) := fun X Y ↦ by simp only [Adjunction.rightAdjointLaxMonoidal_μ, Adjunction.homEquiv_unit] diff --git a/Mathlib/CategoryTheory/Monoidal/Mon_.lean b/Mathlib/CategoryTheory/Monoidal/Mon_.lean index 5669525f6ab4e..cff040d32fd90 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mon_.lean @@ -519,10 +519,10 @@ instance monMonoidalStruct : MonoidalCategoryStruct (Mon_ C) := tensorObj _ _ ⟶ tensorObj _ _ := { hom := f.hom ⊗ g.hom one_hom := by - dsimp + dsimp [tensorObj] slice_lhs 2 3 => rw [← tensor_comp, Hom.one_hom f, Hom.one_hom g] mul_hom := by - dsimp + dsimp [tensorObj] slice_rhs 1 2 => rw [tensorμ_natural] slice_lhs 2 3 => rw [← tensor_comp, Hom.mul_hom f, Hom.mul_hom g, tensor_comp] simp only [Category.assoc] } diff --git a/Mathlib/CategoryTheory/MorphismProperty/Limits.lean b/Mathlib/CategoryTheory/MorphismProperty/Limits.lean index 053633f2b164d..b1bfa27cc46f1 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/Limits.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/Limits.lean @@ -252,7 +252,7 @@ lemma IsStableUnderProductsOfShape.mk (J : Type*) (IsLimit.conePointUniqueUpToIso hc₂ (limit.isLimit X₂) ≪≫ (Pi.isoLimit _).symm) ?_ apply limit.hom_ext rintro ⟨j⟩ - simp + simp [φ] /-- The condition that a property of morphisms is stable by finite products. -/ class IsStableUnderFiniteProducts : Prop where diff --git a/Mathlib/CategoryTheory/Quotient/Linear.lean b/Mathlib/CategoryTheory/Quotient/Linear.lean index bdb9c5b50e4a7..492d4c887729f 100644 --- a/Mathlib/CategoryTheory/Quotient/Linear.lean +++ b/Mathlib/CategoryTheory/Quotient/Linear.lean @@ -49,30 +49,30 @@ lemma smul_eq (hr : ∀ (a : R) ⦃X Y : C⦄ (f₁ f₂ : X ⟶ Y) (_ : r f₁ def module' (hr : ∀ (a : R) ⦃X Y : C⦄ (f₁ f₂ : X ⟶ Y) (_ : r f₁ f₂), r (a • f₁) (a • f₂)) [Preadditive (Quotient r)] [(functor r).Additive] (X Y : C) : Module R ((functor r).obj X ⟶ (functor r).obj Y) := - letI := smul r hr ((functor r).obj X) ((functor r).obj Y) + letI smul := smul r hr ((functor r).obj X) ((functor r).obj Y) { smul_zero := fun a => by dsimp rw [← (functor r).map_zero X Y, smul_eq, smul_zero] zero_smul := fun f => by obtain ⟨f, rfl⟩ := (functor r).map_surjective f - dsimp + dsimp [smul] rw [zero_smul, Functor.map_zero] one_smul := fun f => by obtain ⟨f, rfl⟩ := (functor r).map_surjective f - dsimp + dsimp [smul] rw [one_smul] mul_smul := fun a b f => by obtain ⟨f, rfl⟩ := (functor r).map_surjective f - dsimp + dsimp [smul] rw [mul_smul] smul_add := fun a f g => by obtain ⟨f, rfl⟩ := (functor r).map_surjective f obtain ⟨g, rfl⟩ := (functor r).map_surjective g - dsimp + dsimp [smul] rw [← (functor r).map_add, smul_eq, ← (functor r).map_add, smul_add] add_smul := fun a b f => by obtain ⟨f, rfl⟩ := (functor r).map_surjective f - dsimp + dsimp [smul] rw [add_smul, Functor.map_add] } /-- Auxiliary definition for `Quotient.linear`. -/ diff --git a/Mathlib/CategoryTheory/Sites/Coherent/Comparison.lean b/Mathlib/CategoryTheory/Sites/Coherent/Comparison.lean index f0f439fb8e122..de848d9cbf352 100644 --- a/Mathlib/CategoryTheory/Sites/Coherent/Comparison.lean +++ b/Mathlib/CategoryTheory/Sites/Coherent/Comparison.lean @@ -45,11 +45,11 @@ instance [FinitaryPreExtensive C] [Preregular C] : Precoherent C where let π' := fun a ↦ pullback.fst g' (Sigma.ι X₁ a) have _ := FinitaryPreExtensive.sigma_desc_iso (fun a ↦ Sigma.ι X₁ a) g' inferInstance refine ⟨X₂, π₂, ?_, ?_⟩ - · have : (Sigma.desc π' ≫ g) = Sigma.desc π₂ := by ext; simp + · have : (Sigma.desc π' ≫ g) = Sigma.desc π₂ := by ext; simp [π₂, π'] rw [← effectiveEpi_desc_iff_effectiveEpiFamily, ← this] infer_instance · refine ⟨id, fun b ↦ pullback.snd _ _, fun b ↦ ?_⟩ - simp only [π₂, id_eq, Category.assoc, ← hg] + simp only [X₂, π₂, id_eq, Category.assoc, ← hg] rw [← Category.assoc, pullback.condition] simp diff --git a/Mathlib/CategoryTheory/Sites/CoverLifting.lean b/Mathlib/CategoryTheory/Sites/CoverLifting.lean index ff75491a36dbb..631707e2b37e5 100644 --- a/Mathlib/CategoryTheory/Sites/CoverLifting.lean +++ b/Mathlib/CategoryTheory/Sites/CoverLifting.lean @@ -145,7 +145,7 @@ lemma liftAux_map {Y : C} (f : G.obj Y ⟶ X) {W : C} (g : W ⟶ Y) (i : S.Arrow { g₁ := 𝟙 _ g₂ := h w := by simpa using w.symm } - simpa using s.condition r ) + simpa [r] using s.condition r ) lemma liftAux_map' {Y Y' : C} (f : G.obj Y ⟶ X) (f' : G.obj Y' ⟶ X) {W : C} (a : W ⟶ Y) (b : W ⟶ Y') (w : G.map a ≫ f = G.map b ≫ f') : diff --git a/Mathlib/CategoryTheory/Sites/CoverPreserving.lean b/Mathlib/CategoryTheory/Sites/CoverPreserving.lean index 67fded0c9b4c2..2e0f39951144f 100644 --- a/Mathlib/CategoryTheory/Sites/CoverPreserving.lean +++ b/Mathlib/CategoryTheory/Sites/CoverPreserving.lean @@ -132,7 +132,7 @@ theorem compatiblePreservingOfFlat {C : Type u₁} [Category.{v₁} C] {D : Type simp conv_lhs => rw [eq₁] conv_rhs => rw [eq₂] - simp only [op_comp, Functor.map_comp, types_comp_apply, eqToHom_op, eqToHom_map] + simp only [c, op_comp, Functor.map_comp, types_comp_apply, eqToHom_op, eqToHom_map] apply congr_arg -- Porting note: was `congr 1` which for some reason doesn't do anything here -- despite goal being of the form f a = f b, with f=`ℱ.val.map (Quiver.Hom.op c'.pt.hom)` /- diff --git a/Mathlib/CategoryTheory/Sites/DenseSubsite/SheafEquiv.lean b/Mathlib/CategoryTheory/Sites/DenseSubsite/SheafEquiv.lean index 7cc5b9aefa271..487e2c24802c2 100644 --- a/Mathlib/CategoryTheory/Sites/DenseSubsite/SheafEquiv.lean +++ b/Mathlib/CategoryTheory/Sites/DenseSubsite/SheafEquiv.lean @@ -82,8 +82,9 @@ lemma isIso_ranCounit_app_of_isDenseSubsite (Y : Sheaf J A) (U X) : ⟨_, I.f ≫ i, ⟨l _ _ _ _ _ I.hf, by simp [hl]⟩⟩ refine Eq.trans ?_ (Y.2.amalgamate_map _ _ _ I').symm apply (Y.2 X _ (IsDenseSubsite.equalizer_mem J K G (l _ _ _ _ _ I.hf) - (l _ _ _ _ _ I'.hf) (by simp [hl]))).isSeparatedFor.ext fun V iUV (hiUV : _ = _) ↦ ?_ - simp [← Functor.map_comp, ← op_comp, hiUV] + (l _ _ _ _ _ I'.hf) (by simp [I', hl]))).isSeparatedFor.ext + fun V iUV (hiUV : _ = _) ↦ ?_ + simp [I', ← Functor.map_comp, ← op_comp, hiUV] refine ⟨(isPointwiseRightKanExtensionRanCounit G.op Y.1 (.op (G.obj U))).lift c, ?_⟩ · have := (isPointwiseRightKanExtensionRanCounit G.op Y.1 (.op (G.obj U))).fac c (.mk (𝟙 _)) simp only [id_obj, comp_obj, StructuredArrow.proj_obj, StructuredArrow.mk_right, diff --git a/Mathlib/CategoryTheory/Sites/EffectiveEpimorphic.lean b/Mathlib/CategoryTheory/Sites/EffectiveEpimorphic.lean index f0f3a969c4866..66852416acf33 100644 --- a/Mathlib/CategoryTheory/Sites/EffectiveEpimorphic.lean +++ b/Mathlib/CategoryTheory/Sites/EffectiveEpimorphic.lean @@ -68,9 +68,9 @@ def isColimitOfEffectiveEpiStruct {X Y : C} (f : Y ⟶ X) (Hf : EffectiveEpiStru let Y' : D := ⟨Over.mk f, 𝟙 _, by simp⟩ let Z' : D := ⟨Over.mk (g₁ ≫ f), g₁, rfl⟩ let g₁' : Z' ⟶ Y' := Over.homMk g₁ - let g₂' : Z' ⟶ Y' := Over.homMk g₂ (by simp [h]) + let g₂' : Z' ⟶ Y' := Over.homMk g₂ (by simp [Y', Z', h]) change F.map g₁' ≫ _ = F.map g₂' ≫ _ - simp only [S.w] + simp only [Y', F, S.w] fac := by rintro S ⟨T,g,hT⟩ dsimp @@ -113,7 +113,7 @@ def effectiveEpiStructOfIsColimit {X Y : C} (f : Y ⟶ X) intro W e h dsimp have := Hf.fac (aux e h) ⟨Over.mk f, 𝟙 _, by simp⟩ - dsimp at this; rw [this]; clear this + dsimp [aux] at this; rw [this]; clear this nth_rewrite 2 [← Category.id_comp e] apply h generalize_proofs hh @@ -181,7 +181,7 @@ def isColimitOfEffectiveEpiFamilyStruct {B : C} {α : Type*} let i₁ : Z' ⟶ A₁ := Over.homMk g₁ let i₂ : Z' ⟶ A₂ := Over.homMk g₂ change F.map i₁ ≫ _ = F.map i₂ ≫ _ - simp only [S.w] + simp only [F, A₁, A₂, S.w] fac := by intro S ⟨T, a, (g : T.left ⟶ X a), hT⟩ dsimp @@ -226,7 +226,7 @@ def effectiveEpiFamilyStructOfIsColimit {B : C} {α : Type*} intro W e h a dsimp have := H.fac (aux e h) ⟨Over.mk (π a), a, 𝟙 _, by simp⟩ - dsimp at this; rw [this]; clear this + dsimp [aux] at this; rw [this]; clear this conv_rhs => rw [← Category.id_comp (e a)] apply h generalize_proofs h1 h2 diff --git a/Mathlib/CategoryTheory/Sites/Over.lean b/Mathlib/CategoryTheory/Sites/Over.lean index 2b06a21100f61..0714d53204dde 100644 --- a/Mathlib/CategoryTheory/Sites/Over.lean +++ b/Mathlib/CategoryTheory/Sites/Over.lean @@ -84,7 +84,7 @@ lemma overEquiv_pullback {X : C} {Y₁ Y₂ : Over X} (f : Y₁ ⟶ Y₂) (S : S let T := Over.mk (b ≫ W.hom) let c : T ⟶ Y₁ := Over.homMk g (by dsimp [T]; rw [← Over.w a, ← reassoc_of% w, Over.w f]) let d : T ⟶ W := Over.homMk b - refine ⟨T, c, 𝟙 Z, ?_, by simp [c]⟩ + refine ⟨T, c, 𝟙 Z, ?_, by simp [T, c]⟩ rw [show c ≫ f = d ≫ a by ext; exact w] exact S.downward_closed h _ diff --git a/Mathlib/CategoryTheory/Triangulated/Functor.lean b/Mathlib/CategoryTheory/Triangulated/Functor.lean index 884a5b5ce77dc..2bafc57433276 100644 --- a/Mathlib/CategoryTheory/Triangulated/Functor.lean +++ b/Mathlib/CategoryTheory/Triangulated/Functor.lean @@ -194,7 +194,7 @@ noncomputable instance [F.IsTriangulated] : comm₃ := by simp } exact isIso₂_of_isIso₁₃ φ (F.map_distinguished _ (binaryProductTriangle_distinguished X₁ X₃)) (binaryProductTriangle_distinguished _ _) - (by dsimp; infer_instance) (by dsimp; infer_instance) + (by dsimp [φ]; infer_instance) (by dsimp [φ]; infer_instance) instance (priority := 100) [F.IsTriangulated] : F.Additive := F.additive_of_preserves_binary_products diff --git a/Mathlib/Combinatorics/Additive/Energy.lean b/Mathlib/Combinatorics/Additive/Energy.lean index 2cd35ab13760c..e1f1d36d3e389 100644 --- a/Mathlib/Combinatorics/Additive/Energy.lean +++ b/Mathlib/Combinatorics/Additive/Energy.lean @@ -178,7 +178,7 @@ lemma mulEnergy_univ_left : Eₘ[univ, t] = Fintype.card α * t.card ^ 2 := by let f : α × α × α → (α × α) × α × α := fun x => ((x.1 * x.2.2, x.1 * x.2.1), x.2) have : (↑((univ : Finset α) ×ˢ t ×ˢ t) : Set (α × α × α)).InjOn f := by rintro ⟨a₁, b₁, c₁⟩ _ ⟨a₂, b₂, c₂⟩ h₂ h - simp_rw [Prod.ext_iff] at h + simp_rw [f, Prod.ext_iff] at h obtain ⟨h, rfl, rfl⟩ := h rw [mul_right_cancel h.1] rw [← card_image_of_injOn this] @@ -187,7 +187,7 @@ lemma mulEnergy_univ_left : Eₘ[univ, t] = Fintype.card α * t.card ^ 2 := by Prod.exists] refine ⟨fun h => ⟨a.1.1 * a.2.2⁻¹, _, _, h.1, by simp [f, mul_right_comm, h.2]⟩, ?_⟩ rintro ⟨b, c, d, hcd, rfl⟩ - simpa [mul_right_comm] + simpa [f, mul_right_comm] @[to_additive (attr := simp)] lemma mulEnergy_univ_right : Eₘ[s, univ] = Fintype.card α * s.card ^ 2 := by diff --git a/Mathlib/Combinatorics/Hall/Basic.lean b/Mathlib/Combinatorics/Hall/Basic.lean index b7de187e22a39..1cd190b29bab4 100644 --- a/Mathlib/Combinatorics/Hall/Basic.lean +++ b/Mathlib/Combinatorics/Hall/Basic.lean @@ -184,10 +184,11 @@ theorem Fintype.all_card_le_rel_image_card_iff_exists_injective {α : Type u} { rw [← Set.toFinset_card] apply congr_arg ext b - -- Porting note: added `Set.mem_toFinset` (fixed by lean#6123) - simp [Rel.image, (Set.mem_toFinset)] - -- Porting note: added `Set.mem_toFinset` (fixed by lean#6123) - have h' : ∀ (f : α → β) (x), r x (f x) ↔ f x ∈ r' x := by simp [Rel.image, (Set.mem_toFinset)] + -- Porting note: added `Set.mem_toFinset` (fixed by https://github.com/leanprover/lean4/pull/6123?) + simp [r', Rel.image, (Set.mem_toFinset)] + -- Porting note: added `Set.mem_toFinset` (fixed by https://github.com/leanprover/lean4/pull/6123?) + have h' : ∀ (f : α → β) (x), r x (f x) ↔ f x ∈ r' x := by + simp [r', Rel.image, (Set.mem_toFinset)] simp only [h, h'] apply Finset.all_card_le_biUnion_card_iff_exists_injective diff --git a/Mathlib/Combinatorics/Hall/Finite.lean b/Mathlib/Combinatorics/Hall/Finite.lean index 658e17aae213d..73906569460be 100644 --- a/Mathlib/Combinatorics/Hall/Finite.lean +++ b/Mathlib/Combinatorics/Hall/Finite.lean @@ -109,7 +109,8 @@ theorem hall_hard_inductive_step_A {n : ℕ} (hn : Fintype.card ι = n + 1) have key : ∀ {x}, y ≠ f' x := by intro x h simpa [t', ← h] using hfr x - by_cases h₁ : z₁ = x <;> by_cases h₂ : z₂ = x <;> simp [h₁, h₂, hfinj.eq_iff, key, key.symm] + by_cases h₁ : z₁ = x <;> by_cases h₂ : z₂ = x <;> + simp [h₁, h₂, hfinj.eq_iff, key, key.symm] · intro z simp only [ne_eq, Set.mem_setOf_eq] split_ifs with hz diff --git a/Mathlib/Data/Finite/Card.lean b/Mathlib/Data/Finite/Card.lean index 00700733c3af1..a1193eb881ff5 100644 --- a/Mathlib/Data/Finite/Card.lean +++ b/Mathlib/Data/Finite/Card.lean @@ -46,7 +46,7 @@ theorem Nat.card_eq (α : Type*) : Nat.card α = if _ : Finite α then @Fintype.card α (Fintype.ofFinite α) else 0 := by cases finite_or_infinite α · letI := Fintype.ofFinite α - simp only [*, Nat.card_eq_fintype_card, dif_pos] + simp only [this, *, Nat.card_eq_fintype_card, dif_pos] · simp only [*, card_eq_zero_of_infinite, not_finite_iff_infinite.mpr, dite_false] theorem Finite.card_pos_iff [Finite α] : 0 < Nat.card α ↔ Nonempty α := by diff --git a/Mathlib/Data/Fintype/Card.lean b/Mathlib/Data/Fintype/Card.lean index 22e95987df0fd..d9cf21d019a1f 100644 --- a/Mathlib/Data/Fintype/Card.lean +++ b/Mathlib/Data/Fintype/Card.lean @@ -1114,7 +1114,7 @@ theorem List.exists_pw_disjoint_with_card {α : Type*} [Fintype α] intro a ha conv_rhs => rw [← List.map_id a] rw [List.map_pmap] - simp only [Fin.valEmbedding_apply, Fin.val_mk, List.pmap_eq_map, List.map_id', List.map_id] + simp [klift, Fin.valEmbedding_apply, Fin.val_mk, List.pmap_eq_map, List.map_id'] use l.map (List.map (Fintype.equivFin α).symm) constructor · -- length diff --git a/Mathlib/Data/List/Indexes.lean b/Mathlib/Data/List/Indexes.lean index af5d70ec012f2..dfa2e8e7d9973 100644 --- a/Mathlib/Data/List/Indexes.lean +++ b/Mathlib/Data/List/Indexes.lean @@ -286,9 +286,8 @@ theorem mapIdxMAux'_eq_mapIdxMGo {α} (f : ℕ → α → m PUnit) (as : List α · simp only [mapIdxMAux', seqRight_eq, map_eq_pure_bind, seq_eq_bind, bind_pure_unit, LawfulMonad.bind_assoc, pure_bind, mapIdxM.go, seq_pure] generalize (f (Array.size arr) head) = head - let arr_1 := arr.push ⟨⟩ - have : arr_1.size = arr.size + 1 := Array.size_push arr ⟨⟩ - rw [← this, ih arr_1] + have : (arr.push ⟨⟩).size = arr.size + 1 := Array.size_push arr ⟨⟩ + rw [← this, ih] simp only [seqRight_eq, map_eq_pure_bind, seq_pure, LawfulMonad.bind_assoc, pure_bind] theorem mapIdxM'_eq_mapIdxM {α} (f : ℕ → α → m PUnit) (as : List α) : diff --git a/Mathlib/Data/Real/Pi/Irrational.lean b/Mathlib/Data/Real/Pi/Irrational.lean index afebb88a105b2..34d2652435e86 100644 --- a/Mathlib/Data/Real/Pi/Irrational.lean +++ b/Mathlib/Data/Real/Pi/Irrational.lean @@ -91,7 +91,7 @@ private lemma recursion' (n : ℕ) : 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] + · simp_rw [u₁, v₁', f, ← 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] diff --git a/Mathlib/Data/Set/Card.lean b/Mathlib/Data/Set/Card.lean index 4588229a6887f..56ac372f3623f 100644 --- a/Mathlib/Data/Set/Card.lean +++ b/Mathlib/Data/Set/Card.lean @@ -762,8 +762,8 @@ theorem surj_on_of_inj_on_of_ncard_le {t : Set β} (f : ∀ a ∈ s, β) (hf : have hft' := Fintype.ofInjective f' finj set f'' : ∀ a, a ∈ s.toFinset → β := fun a h ↦ f a (by simpa using h) convert @Finset.surj_on_of_inj_on_of_card_le _ _ _ t.toFinset f'' _ _ _ _ (by simpa) using 1 - · simp - · simp [hf] + · simp [f''] + · simp [f'', hf] · intros a₁ a₂ ha₁ ha₂ h rw [mem_toFinset] at ha₁ ha₂ exact hinj _ _ ha₁ ha₂ h diff --git a/Mathlib/Data/Setoid/Partition.lean b/Mathlib/Data/Setoid/Partition.lean index 1aa847fee7ff8..ddc6a2e70e0d9 100644 --- a/Mathlib/Data/Setoid/Partition.lean +++ b/Mathlib/Data/Setoid/Partition.lean @@ -166,7 +166,7 @@ noncomputable def quotientEquivClasses (r : Setoid α) : Quotient r ≃ Setoid.c · intro (q_a : Quotient r) (q_b : Quotient r) h_eq induction' q_a using Quotient.ind with a induction' q_b using Quotient.ind with b - simp only [Subtype.ext_iff, Quotient.lift_mk, Subtype.ext_iff] at h_eq + simp only [f, Quotient.lift_mk, Subtype.ext_iff] at h_eq apply Quotient.sound show a ∈ { x | r x b } rw [← h_eq] diff --git a/Mathlib/FieldTheory/Extension.lean b/Mathlib/FieldTheory/Extension.lean index 142e5ed48085b..131f0f917c5ff 100644 --- a/Mathlib/FieldTheory/Extension.lean +++ b/Mathlib/FieldTheory/Extension.lean @@ -183,7 +183,7 @@ theorem nonempty_algHom_of_exist_lifts_finset [alg : Algebra.IsAlgebraic F E] have := finiteDimensional_adjoin (S := {α}) fun _ _ ↦ ((alg.tower_top ϕ.carrier).isIntegral).1 _ let L (σ : Λ) : Lifts F E K := ⟨ϕ.carrier⟮α⟯.restrictScalars F, σ.restrictScalars F⟩ have hL (σ : Λ) : ϕ < L σ := lt_iff.mpr - ⟨by simpa only [restrictScalars_adjoin_eq_sup, left_lt_sup, adjoin_simple_le_iff], + ⟨by simpa only [L, restrictScalars_adjoin_eq_sup, left_lt_sup, adjoin_simple_le_iff], AlgHom.coe_ringHom_injective σ.comp_algebraMap⟩ have ⟨(ϕ_ext : ϕ.IsExtendible), ϕ_max⟩ := maximal_iff_forall_gt.mp hϕ simp_rw [Set.mem_setOf, IsExtendible] at ϕ_max; push_neg at ϕ_max diff --git a/Mathlib/Geometry/Euclidean/Angle/Oriented/Affine.lean b/Mathlib/Geometry/Euclidean/Angle/Oriented/Affine.lean index eb38b5464c6a8..5d6360856ed2d 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Oriented/Affine.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Oriented/Affine.lean @@ -46,8 +46,8 @@ def oangle (p₁ p₂ p₃ : P) : Real.Angle := theorem continuousAt_oangle {x : P × P × P} (hx12 : x.1 ≠ x.2.1) (hx32 : x.2.2 ≠ x.2.1) : ContinuousAt (fun y : P × P × P => ∡ y.1 y.2.1 y.2.2) x := by let f : P × P × P → V × V := fun y => (y.1 -ᵥ y.2.1, y.2.2 -ᵥ y.2.1) - have hf1 : (f x).1 ≠ 0 := by simp [hx12] - have hf2 : (f x).2 ≠ 0 := by simp [hx32] + have hf1 : (f x).1 ≠ 0 := by simp [f, hx12] + have hf2 : (f x).2 ≠ 0 := by simp [f, hx32] exact (o.continuousAt_oangle hf1 hf2).comp ((continuous_fst.vsub continuous_snd.fst).prod_mk (continuous_snd.snd.vsub continuous_snd.fst)).continuousAt diff --git a/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean b/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean index 36b42892368e8..782d2803dcdb5 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean @@ -47,8 +47,8 @@ nonrec def angle (p1 p2 p3 : P) : ℝ := theorem continuousAt_angle {x : P × P × P} (hx12 : x.1 ≠ x.2.1) (hx32 : x.2.2 ≠ x.2.1) : ContinuousAt (fun y : P × P × P => ∠ y.1 y.2.1 y.2.2) x := by let f : P × P × P → V × V := fun y => (y.1 -ᵥ y.2.1, y.2.2 -ᵥ y.2.1) - have hf1 : (f x).1 ≠ 0 := by simp [hx12] - have hf2 : (f x).2 ≠ 0 := by simp [hx32] + have hf1 : (f x).1 ≠ 0 := by simp [f, hx12] + have hf2 : (f x).2 ≠ 0 := by simp [f, hx32] exact (InnerProductGeometry.continuousAt_angle hf1 hf2).comp ((continuous_fst.vsub continuous_snd.fst).prod_mk (continuous_snd.snd.vsub continuous_snd.fst)).continuousAt diff --git a/Mathlib/Geometry/Euclidean/Circumcenter.lean b/Mathlib/Geometry/Euclidean/Circumcenter.lean index dd63de61c9efd..a2bd062cbe191 100644 --- a/Mathlib/Geometry/Euclidean/Circumcenter.lean +++ b/Mathlib/Geometry/Euclidean/Circumcenter.lean @@ -645,7 +645,7 @@ theorem reflection_circumcenter_eq_affineCombination_of_pointsWithCircumcenter { (orthogonalProjection W s.circumcenter : P) = ↑((s.face hc).orthogonalProjectionSpan s.circumcenter) := by apply eq_orthogonalProjection_of_eq_subspace - simp + simp [W] rw [EuclideanGeometry.reflection_apply, h_faces, s.orthogonalProjection_circumcenter hc, circumcenter_eq_centroid, s.face_centroid_eq_centroid hc, centroid_eq_affineCombination_of_pointsWithCircumcenter, diff --git a/Mathlib/Geometry/Euclidean/MongePoint.lean b/Mathlib/Geometry/Euclidean/MongePoint.lean index 7a59b661c7066..9800925e6fb48 100644 --- a/Mathlib/Geometry/Euclidean/MongePoint.lean +++ b/Mathlib/Geometry/Euclidean/MongePoint.lean @@ -220,9 +220,9 @@ theorem inner_mongePoint_vsub_face_centroid_vsub {n : ℕ} (s : Simplex ℝ P (n · simp_rw [sum_pointsWithCircumcenter, pointsWithCircumcenter_eq_circumcenter, pointsWithCircumcenter_point, Pi.sub_apply, pointWeightsWithCircumcenter] rw [← sum_subset fs.subset_univ _] - · simp_rw [sum_insert (not_mem_singleton.2 h), sum_singleton] + · simp_rw [fs, sum_insert (not_mem_singleton.2 h), sum_singleton] repeat rw [← sum_subset fs.subset_univ _] - · simp_rw [sum_insert (not_mem_singleton.2 h), sum_singleton] + · simp_rw [fs, sum_insert (not_mem_singleton.2 h), sum_singleton] simp [h, Ne.symm h, dist_comm (s.points i₁)] all_goals intro i _ hi; simp [hfs i hi] · intro i _ hi diff --git a/Mathlib/GroupTheory/Exponent.lean b/Mathlib/GroupTheory/Exponent.lean index ea64a5ce67e5e..8adffc6d68bd0 100644 --- a/Mathlib/GroupTheory/Exponent.lean +++ b/Mathlib/GroupTheory/Exponent.lean @@ -439,7 +439,7 @@ theorem exists_orderOf_eq_exponent (hG : ExponentExists G) : ∃ g : G, orderOf nth_rw 1 [← pow_one p] have : 1 = (Nat.factorization (orderOf (t ^ p ^ k))) p + 1 := by rw [hpk', Nat.factorization_div hpk] - simp [hp] + simp [k, hp] rw [this] -- Porting note: convert made to_additive complain apply Nat.pow_succ_factorization_not_dvd (hG.orderOf_pos <| t ^ p ^ k).ne' hp diff --git a/Mathlib/GroupTheory/Goursat.lean b/Mathlib/GroupTheory/Goursat.lean index 3864263abddfd..d24ea2e6457e4 100644 --- a/Mathlib/GroupTheory/Goursat.lean +++ b/Mathlib/GroupTheory/Goursat.lean @@ -152,20 +152,21 @@ lemma goursat : · ext ⟨g, h⟩ constructor · intro hgh - simpa only [mem_map, MonoidHom.mem_range, MonoidHom.prod_apply, Subtype.exists, Prod.exists, - MonoidHom.coe_prodMap, coeSubtype, Prod.mk.injEq, Prod.map_apply, MonoidHom.coe_snd, - exists_eq_right, exists_and_right, exists_eq_right_right, MonoidHom.coe_fst] + simpa only [G', H', mem_map, MonoidHom.mem_range, MonoidHom.prod_apply, Subtype.exists, + Prod.exists, MonoidHom.coe_prodMap, coeSubtype, Prod.mk.injEq, Prod.map_apply, + MonoidHom.coe_snd, exists_eq_right, exists_and_right, exists_eq_right_right, + MonoidHom.coe_fst] using ⟨⟨h, hgh⟩, ⟨g, hgh⟩, g, h, hgh, ⟨rfl, rfl⟩⟩ - · simp only [mem_map, MonoidHom.mem_range, MonoidHom.prod_apply, Subtype.exists, Prod.exists, - MonoidHom.coe_prodMap, coeSubtype, Prod.mk.injEq, Prod.map_apply, MonoidHom.coe_snd, - exists_eq_right, exists_and_right, exists_eq_right_right, MonoidHom.coe_fst, - forall_exists_index, and_imp] + · simp only [G', H', mem_map, MonoidHom.mem_range, MonoidHom.prod_apply, Subtype.exists, + Prod.exists, MonoidHom.coe_prodMap, coeSubtype, Prod.mk.injEq, Prod.map_apply, + MonoidHom.coe_snd, exists_eq_right, exists_and_right, exists_eq_right_right, + MonoidHom.coe_fst, forall_exists_index, and_imp] rintro h₁ hgh₁ g₁ hg₁h g₂ h₂ hg₂h₂ hP hQ simp only [Subtype.ext_iff] at hP hQ rwa [← hP, ← hQ] · convert goursatFst_prod_goursatSnd_le (P.prod Q).range ext ⟨g, h⟩ - simp_rw [MonoidHom.mem_ker, MonoidHom.coe_prodMap, Prod.map_apply, Subgroup.mem_prod, + simp_rw [G', H', MonoidHom.mem_ker, MonoidHom.coe_prodMap, Prod.map_apply, Subgroup.mem_prod, Prod.one_eq_mk, Prod.ext_iff, ← MonoidHom.mem_ker, QuotientGroup.ker_mk'] end Subgroup diff --git a/Mathlib/GroupTheory/HNNExtension.lean b/Mathlib/GroupTheory/HNNExtension.lean index 2c932f875279a..79dab5bd2e7c8 100644 --- a/Mathlib/GroupTheory/HNNExtension.lean +++ b/Mathlib/GroupTheory/HNNExtension.lean @@ -679,8 +679,8 @@ theorem toList_eq_nil_of_mem_of_range (w : ReducedWord G A B) w.toList = [] := by rcases hw with ⟨g, hg⟩ let w' : ReducedWord G A B := { ReducedWord.empty G A B with head := g } - have : w.prod φ = w'.prod φ := by simp [ReducedWord.prod, hg] - simpa using (map_fst_eq_and_of_prod_eq φ this).1 + have : w.prod φ = w'.prod φ := by simp [w', ReducedWord.prod, hg] + simpa [w'] using (map_fst_eq_and_of_prod_eq φ this).1 end ReducedWord diff --git a/Mathlib/GroupTheory/PushoutI.lean b/Mathlib/GroupTheory/PushoutI.lean index 0f672f3da1e25..c275a0b83ee72 100644 --- a/Mathlib/GroupTheory/PushoutI.lean +++ b/Mathlib/GroupTheory/PushoutI.lean @@ -425,7 +425,7 @@ noncomputable def equivPair (i) : NormalWord d ≃ Pair d i := exact w.normalized _ _ (Word.mem_of_mem_equivPair_tail _ hg) } haveI leftInv : Function.LeftInverse (rcons i) toFun := fun w => ext_smul i <| by - simp only [rcons, Word.equivPair_symm, + simp only [toFun, rcons, Word.equivPair_symm, Word.equivPair_smul_same, Word.equivPair_tail_eq_inv_smul, Word.rcons_eq_smul, MonoidHom.apply_ofInjective_symm, equiv_fst_eq_mul_inv, mul_assoc, map_mul, map_inv, mul_smul, inv_smul_smul, smul_inv_smul] @@ -679,11 +679,12 @@ theorem inf_of_range_eq_base_range exact hx (of_apply_eq_base φ j y ▸ MonoidHom.mem_range.2 ⟨y, rfl⟩) let w : Word G := ⟨[⟨_, g₁⟩, ⟨_, g₂⁻¹⟩], by simp_all, by simp_all⟩ have hw : Reduced φ w := by - simp only [not_exists, ne_eq, Reduced, List.find?, List.mem_cons, List.mem_singleton, - forall_eq_or_imp, not_false_eq_true, forall_const, forall_eq, true_and, hg₁r, hg₂r, - List.mem_nil_iff, false_imp_iff, imp_true_iff, and_true, inv_mem_iff] + simp only [w, not_exists, ne_eq, Reduced, List.find?, List.mem_cons, + List.mem_singleton, forall_eq_or_imp, not_false_eq_true, forall_const, forall_eq, + true_and, hg₁r, hg₂r, List.mem_nil_iff, false_imp_iff, imp_true_iff, and_true, + inv_mem_iff] have := hw.eq_empty_of_mem_range hφ (by - simp only [Word.prod, List.map_cons, List.prod_cons, List.prod_nil, + simp only [w, Word.prod, List.map_cons, List.prod_cons, List.prod_nil, List.map_nil, map_mul, ofCoprodI_of, hg₁, hg₂, map_inv, map_one, mul_one, mul_inv_cancel, one_mem]) simp [w, Word.empty] at this) diff --git a/Mathlib/LinearAlgebra/AffineSpace/Combination.lean b/Mathlib/LinearAlgebra/AffineSpace/Combination.lean index 7170f542e27e4..25dfb8d294ad8 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Combination.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Combination.lean @@ -412,9 +412,9 @@ theorem attach_affineCombination_of_injective [DecidableEq P] (s : Finset P) (w change univ.sum g₁ = (image f univ).sum g₂ have hgf : g₁ = g₂ ∘ f := by ext - simp + simp [g₁, g₂] rw [hgf, sum_image] - · simp only [Function.comp_apply] + · simp only [g₁, g₂,Function.comp_apply] · exact fun _ _ _ _ hxy => hf hxy theorem attach_affineCombination_coe (s : Finset P) (w : P → k) : @@ -927,7 +927,7 @@ theorem affineCombination_mem_affineSpan [Nontrivial k] {s : Finset ι} {w : ι cases' hn with i1 hi1 let w1 : ι → k := Function.update (Function.const ι 0) i1 1 have hw1 : ∑ i ∈ s, w1 i = 1 := by - simp only [Function.const_zero, Finset.sum_update_of_mem hi1, Pi.zero_apply, + simp only [w1, Function.const_zero, Finset.sum_update_of_mem hi1, Pi.zero_apply, Finset.sum_const_zero, add_zero] have hw1s : s.affineCombination k p w1 = p i1 := s.affineCombination_of_eq_one_of_eq_zero w1 p hi1 (Function.update_same _ _ _) fun _ _ hne => @@ -1057,8 +1057,8 @@ theorem mem_affineSpan_iff_eq_weightedVSubOfPoint_vadd [Nontrivial k] (p : ι let w' : ι → k := Function.update w j (1 - (s \ {j}).sum w) have h₁ : (insert j s).sum w' = 1 := by by_cases hj : j ∈ s - · simp [Finset.sum_update_of_mem hj, Finset.insert_eq_of_mem hj] - · simp [Finset.sum_insert hj, Finset.sum_update_of_not_mem hj, hj] + · simp [w', Finset.sum_update_of_mem hj, Finset.insert_eq_of_mem hj] + · simp [w', Finset.sum_insert hj, Finset.sum_update_of_not_mem hj, hj] have hww : ∀ i, i ≠ j → w i = w' i := by intro i hij simp [w', hij] diff --git a/Mathlib/LinearAlgebra/AffineSpace/Independent.lean b/Mathlib/LinearAlgebra/AffineSpace/Independent.lean index 308268392107c..1adcc9da63cc8 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Independent.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Independent.lean @@ -457,14 +457,14 @@ theorem exists_nontrivial_relation_sum_zero_of_not_affine_ind {t : Finset V} simp only [Finset.weightedVSub_eq_weightedVSubOfPoint_of_sum_eq_zero _ w ((↑) : t → V) hw 0, vsub_eq_sub, Finset.weightedVSubOfPoint_apply, sub_zero] at hwt let f : ∀ x : V, x ∈ t → k := fun x hx => w ⟨x, hx⟩ - refine ⟨fun x => if hx : x ∈ t then f x hx else (0 : k), ?_, ?_, by use i; simp [hi]⟩ + refine ⟨fun x => if hx : x ∈ t then f x hx else (0 : k), ?_, ?_, by use i; simp [f, hi]⟩ on_goal 1 => suffices (∑ e ∈ t, dite (e ∈ t) (fun hx => f e hx • e) fun _ => 0) = 0 by convert this rename V => x by_cases hx : x ∈ t <;> simp [hx] all_goals - simp only [Finset.sum_dite_of_true fun _ h => h, Finset.mk_coe, hwt, hw] + simp only [f, Finset.sum_dite_of_true fun _ h => h, Finset.mk_coe, hwt, hw] variable {s : Finset ι} {w w₁ w₂ : ι → k} {p : ι → V} @@ -618,7 +618,7 @@ theorem affineIndependent_of_ne {p₁ p₂ : P} (h : p₁ ≠ p₂) : AffineInde ext fin_cases i · simp at hi - · simp + · simp [i₁] haveI : Unique { x // x ≠ (0 : Fin 2) } := ⟨⟨i₁⟩, he'⟩ apply linearIndependent_unique rw [he' default] diff --git a/Mathlib/LinearAlgebra/Charpoly/ToMatrix.lean b/Mathlib/LinearAlgebra/Charpoly/ToMatrix.lean index 6df8a1d820cb3..712406dae7adc 100644 --- a/Mathlib/LinearAlgebra/Charpoly/ToMatrix.lean +++ b/Mathlib/LinearAlgebra/Charpoly/ToMatrix.lean @@ -68,7 +68,8 @@ theorem charpoly_toMatrix {ι : Type w} [DecidableEq ι] [Fintype ι] (b : Basis rw [basis_toMatrix_mul_linearMap_toMatrix_mul_basis_toMatrix] _ = det (scalar ι' X - C.mapMatrix (φ₁ P * φ₂ A' * φ₃ Q)) := by rw [reindexLinearEquiv_mul, reindexLinearEquiv_mul] - _ = det (scalar ι' X - C.mapMatrix (φ₁ P) * C.mapMatrix A' * C.mapMatrix (φ₃ Q)) := by simp [φ₂] + _ = det (scalar ι' X - C.mapMatrix (φ₁ P) * C.mapMatrix A' * C.mapMatrix (φ₃ Q)) := by + simp [φ₁, φ₂, φ₃, ι'] _ = det (scalar ι' X * C.mapMatrix (φ₁ P) * C.mapMatrix (φ₃ Q) - C.mapMatrix (φ₁ P) * C.mapMatrix A' * C.mapMatrix (φ₃ Q)) := by rw [Matrix.mul_assoc ((scalar ι') X), hPQ, Matrix.mul_one] diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/Basic.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/Basic.lean index 3596252e1e410..2a899d045c834 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/Basic.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/Basic.lean @@ -207,7 +207,7 @@ theorem induction {C : CliffordAlgebra Q → Prop} -- the mapping through the subalgebra is the identity have of_id : AlgHom.id R (CliffordAlgebra Q) = s.val.comp (lift Q of) := by ext - simp [of] + simp [of, h] -- Porting note: `simp` can't apply this erw [LinearMap.codRestrict_apply] -- finding a proof is finding an element of the subalgebra diff --git a/Mathlib/LinearAlgebra/Dimension/DivisionRing.lean b/Mathlib/LinearAlgebra/Dimension/DivisionRing.lean index 19bb2b8c796ae..99dc600e166de 100644 --- a/Mathlib/LinearAlgebra/Dimension/DivisionRing.lean +++ b/Mathlib/LinearAlgebra/Dimension/DivisionRing.lean @@ -176,7 +176,7 @@ theorem linearIndependent_iff_card_eq_finrank_span {ι : Type*} [Fintype ι] {b have h : span K (f '' Set.range b') = map f (span K (Set.range b')) := span_image f have hf : f '' Set.range b' = Set.range b := by ext x - simp [f, Set.mem_image, Set.mem_range] + simp [f, b', Set.mem_image, Set.mem_range] rw [hf] at h have hx : (x : V) ∈ span K (Set.range b) := x.property simp_rw [h] at hx diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index 18b5e4ddd8cd1..398f24af5888d 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -803,7 +803,8 @@ theorem _root_.mem_span_of_iInf_ker_le_ker [Finite ι] {L : ι → E →ₗ[𝕜 rw [← p.liftQ_mkQ K h] ext x convert LinearMap.congr_fun hK' (p.mkQ x) - simp only [coeFn_sum, Finset.sum_apply, smul_apply, coe_comp, Function.comp_apply, smul_eq_mul] + simp only [L',coeFn_sum, Finset.sum_apply, smul_apply, coe_comp, Function.comp_apply, + smul_eq_mul] end Submodule diff --git a/Mathlib/LinearAlgebra/FreeModule/IdealQuotient.lean b/Mathlib/LinearAlgebra/FreeModule/IdealQuotient.lean index e4b38dca36171..7e5fc7690e990 100644 --- a/Mathlib/LinearAlgebra/FreeModule/IdealQuotient.lean +++ b/Mathlib/LinearAlgebra/FreeModule/IdealQuotient.lean @@ -41,9 +41,7 @@ noncomputable def quotientEquivPiSpan (I : Ideal S) (b : Basis ι R S) (hI : I have ab_eq := I.selfBasis_def b hI have mem_I_iff : ∀ x, x ∈ I ↔ ∀ i, a i ∣ b'.repr x i := by intro x - -- Porting note: these lines used to be `simp_rw [ab.mem_ideal_iff', ab_eq]` - rw [ab.mem_ideal_iff'] - simp_rw [ab_eq] + simp_rw [ab.mem_ideal_iff', ab, ab_eq] have : ∀ (c : ι → R) (i), b'.repr (∑ j : ι, c j • a j • b' j) i = a i * c i := by intro c i simp only [← MulAction.mul_smul, b'.repr_sum_self, mul_comm] diff --git a/Mathlib/LinearAlgebra/Matrix/Block.lean b/Mathlib/LinearAlgebra/Matrix/Block.lean index 401caa2ef1eff..26de8f3e90de3 100644 --- a/Mathlib/LinearAlgebra/Matrix/Block.lean +++ b/Mathlib/LinearAlgebra/Matrix/Block.lean @@ -344,7 +344,7 @@ theorem blockTriangular_inv_of_blockTriangular [LinearOrder α] [Invertible M] have hb' : image b' univ ⊂ image b univ := by convert image_subtype_univ_ssubset_image_univ k b _ (fun a => a < k) (lt_irrefl _) convert max'_mem (α := α) _ _ - have hij' : b' ⟨j, hij.trans hi⟩ < b' ⟨i, hi⟩ := by simp_rw [hij] - simp [hM.inv_toBlock k, (ih (image b' univ) hb' hA rfl hij').symm] + have hij' : b' ⟨j, hij.trans hi⟩ < b' ⟨i, hi⟩ := by simp_rw [b', hij] + simp [A, hM.inv_toBlock k, (ih (image b' univ) hb' hA rfl hij').symm] end Matrix diff --git a/Mathlib/LinearAlgebra/Matrix/Charpoly/Coeff.lean b/Mathlib/LinearAlgebra/Matrix/Charpoly/Coeff.lean index f136828ae7921..e5d7288e96387 100644 --- a/Mathlib/LinearAlgebra/Matrix/Charpoly/Coeff.lean +++ b/Mathlib/LinearAlgebra/Matrix/Charpoly/Coeff.lean @@ -317,9 +317,9 @@ lemma reverse_charpoly (M : Matrix n n R) : let q : R[T;T⁻¹] := det (1 - scalar n t * M.map LaurentPolynomial.C) have ht : t_inv * t = 1 := by rw [← T_add, neg_add_cancel, T_zero] have hp : toLaurentAlg M.charpoly = p := by - simp [p, charpoly, charmatrix, AlgHom.map_det, map_sub, map_smul'] + simp [p, t, charpoly, charmatrix, AlgHom.map_det, map_sub, map_smul'] have hq : toLaurentAlg M.charpolyRev = q := by - simp [q, charpolyRev, AlgHom.map_det, map_sub, map_smul', smul_eq_diagonal_mul] + simp [q, t, charpolyRev, AlgHom.map_det, map_sub, map_smul', smul_eq_diagonal_mul] suffices t_inv ^ Fintype.card n * p = invert q by apply toLaurent_injective rwa [toLaurent_reverse, ← coe_toLaurentAlg, hp, hq, ← involutive_invert.injective.eq_iff, @@ -327,7 +327,7 @@ lemma reverse_charpoly (M : Matrix n n R) : ← mul_one (Fintype.card n : ℤ), ← T_pow, map_pow, invert_T, mul_comm] rw [← det_smul, smul_sub, scalar_apply, ← diagonal_smul, Pi.smul_def, smul_eq_mul, ht, diagonal_one, invert.map_det] - simp [map_sub, _root_.map_one, _root_.map_mul, t, map_smul', smul_eq_diagonal_mul] + simp [t_inv, map_sub, _root_.map_one, _root_.map_mul, t, map_smul', smul_eq_diagonal_mul] @[simp] lemma eval_charpolyRev : diff --git a/Mathlib/LinearAlgebra/Matrix/Transvection.lean b/Mathlib/LinearAlgebra/Matrix/Transvection.lean index 1c777ee3fb514..9b9593fe097ac 100644 --- a/Mathlib/LinearAlgebra/Matrix/Transvection.lean +++ b/Mathlib/LinearAlgebra/Matrix/Transvection.lean @@ -403,7 +403,7 @@ theorem listTransvecCol_mul_last_col (hM : M (inr unit) (inr unit) ≠ 0) (i : F have A : (listTransvecCol M)[n] = transvection (inl n') (inr unit) (-M (inl n') (inr unit) / M (inr unit) (inr unit)) := by - simp [listTransvecCol] + simp [n', listTransvecCol] simp only [Matrix.mul_assoc, A, List.prod_cons] by_cases h : n' = i · have hni : n = i := by @@ -416,7 +416,7 @@ theorem listTransvecCol_mul_last_col (hM : M (inr unit) (inr unit) ≠ 0) (i : F · have hni : n ≠ i := by rintro rfl cases i - simp at h + simp [n'] at h simp only [ne_eq, inl.injEq, Ne.symm h, not_false_eq_true, transvection_mul_apply_of_ne] rw [IH] rcases le_or_lt (n + 1) i with (hi | hi) @@ -441,7 +441,7 @@ theorem mul_listTransvecRow_last_col_take (i : Fin r ⊕ Unit) {k : ℕ} (hk : k (listTransvecRow M)[k]? = ↑(transvection (inr Unit.unit) (inl k') (-M (inr Unit.unit) (inl k') / M (inr Unit.unit) (inr Unit.unit))) := by - simp only [listTransvecRow, List.ofFnNthVal, hkr, dif_pos, List.getElem?_ofFn] + simp only [k', listTransvecRow, List.ofFnNthVal, hkr, dif_pos, List.getElem?_ofFn] simp only [List.take_succ, ← Matrix.mul_assoc, this, List.prod_append, Matrix.mul_one, List.prod_cons, List.prod_nil, Option.toList_some] rw [mul_transvection_apply_of_ne, IH hkr.le] @@ -476,7 +476,7 @@ theorem mul_listTransvecRow_last_row (hM : M (inr unit) (inr unit) ≠ 0) (i : F (listTransvecRow M)[n]? = ↑(transvection (inr unit) (inl n') (-M (inr unit) (inl n') / M (inr unit) (inr unit))) := by - simp only [listTransvecRow, List.ofFnNthVal, hnr, dif_pos, List.getElem?_ofFn] + simp only [n', listTransvecRow, List.ofFnNthVal, hnr, dif_pos, List.getElem?_ofFn] simp only [List.take_succ, A, ← Matrix.mul_assoc, List.prod_append, Matrix.mul_one, List.prod_cons, List.prod_nil, Option.toList_some] by_cases h : n' = i diff --git a/Mathlib/LinearAlgebra/PID.lean b/Mathlib/LinearAlgebra/PID.lean index 77090d78d128c..9c332cd59b756 100644 --- a/Mathlib/LinearAlgebra/PID.lean +++ b/Mathlib/LinearAlgebra/PID.lean @@ -41,6 +41,6 @@ lemma trace_restrict_eq_of_forall_mem [IsDomain R] [IsPrincipalIdealRing R] contrapose! hi; exact snf.repr_eq_zero_of_nmem_range ⟨_, (hf _)⟩ hi change ∑ i, A i i = ∑ i, B i i rw [← Finset.sum_filter_of_ne (p := fun j ↦ j ∈ Set.range snf.f) (by simpa using aux)] - simp [A] + simp [A, B] end LinearMap diff --git a/Mathlib/LinearAlgebra/PerfectPairing.lean b/Mathlib/LinearAlgebra/PerfectPairing.lean index 46a72d4138489..78fff41b7e142 100644 --- a/Mathlib/LinearAlgebra/PerfectPairing.lean +++ b/Mathlib/LinearAlgebra/PerfectPairing.lean @@ -261,11 +261,11 @@ lemma exists_basis_basis_of_span_eq_top_of_mem_algebraMap suffices span K (Set.range v') = ⊤ by let e := (Module.Finite.finite_basis b).equivFin let b' : Basis _ K M' := Basis.mk hv' (by rw [this]) - exact ⟨_, b.reindex e, b'.reindex e, fun i ↦ by simp [b, b']⟩ + exact ⟨_, b.reindex e, b'.reindex e, fun i ↦ by simp [b, b', v']⟩ suffices span K v = M' by apply Submodule.map_injective_of_injective M'.injective_subtype rw [Submodule.map_span, ← Set.image_univ, Set.image_image] - simpa + simpa [v'] refine le_antisymm (Submodule.span_le.mpr hv₁) fun m hm ↦ ?_ obtain ⟨w, hw₁, hw₂, hw₃⟩ := exists_linearIndependent L (N' : Set N) rw [hN] at hw₂ @@ -282,7 +282,7 @@ lemma exists_basis_basis_of_span_eq_top_of_mem_algebraMap toDualLeft_apply, Basis.dualBasis_repr] exact hp (b j) (by simpa [b] using hv₁ j.2) (bN i) (by simpa [bN] using hw₁ i.2) have hA (i j) : b.toMatrix bM i j ∈ (algebraMap K L).range := - Matrix.mem_subfield_of_mul_eq_one_of_mem_subfield_left e _ (by simp) hp i j + Matrix.mem_subfield_of_mul_eq_one_of_mem_subfield_left e _ (by simp [bM]) hp i j have h_span : span K v = span K (Set.range b) := by simp [b] rw [h_span, Basis.mem_span_iff_repr_mem, ← Basis.toMatrix_mulVec_repr bM b m] exact fun i ↦ Subring.sum_mem _ fun j _ ↦ Subring.mul_mem _ (hA i j) (hj j) diff --git a/Mathlib/LinearAlgebra/Semisimple.lean b/Mathlib/LinearAlgebra/Semisimple.lean index 9f91ee1b94849..f96061725cfc2 100644 --- a/Mathlib/LinearAlgebra/Semisimple.lean +++ b/Mathlib/LinearAlgebra/Semisimple.lean @@ -297,7 +297,7 @@ theorem IsSemisimple.of_mem_adjoin_pair {a : End K M} (ha : a ∈ Algebra.adjoin rotate_left 1 · rw [Ideal.span, ← minpoly.ker_aeval_eq_span_minpoly]; exact id · rintro ⟨p⟩; exact p.induction_on (fun k ↦ by simp [R, Algebra.commute_algebraMap_left]) - (fun p q hp hq ↦ by simpa using hp.add_left hq) + (fun p q hp hq ↦ by simpa [R] using hp.add_left hq) fun n k ↦ by simpa [R, pow_succ, ← mul_assoc _ _ X] using (·.mul_left comm) · simpa only [RingHom.mem_ker, eval₂AlgHom'_apply, eval₂_map, AlgHom.comp_algebraMap_of_tower] using minpoly.aeval K g diff --git a/Mathlib/Logic/Function/Basic.lean b/Mathlib/Logic/Function/Basic.lean index 5f794b63872c6..a1af4bbb91992 100644 --- a/Mathlib/Logic/Function/Basic.lean +++ b/Mathlib/Logic/Function/Basic.lean @@ -267,7 +267,7 @@ theorem not_surjective_Type {α : Type u} (f : α → Type max u v) : ¬Surjecti have hg : Injective g := by intro s t h suffices cast hU (g s).2 = cast hU (g t).2 by - simp only [cast_cast, cast_eq] at this + simp only [g, cast_cast, cast_eq] at this assumption · congr exact cantor_injective g hg diff --git a/Mathlib/MeasureTheory/Constructions/Pi.lean b/Mathlib/MeasureTheory/Constructions/Pi.lean index 670e4e5bf0347..f7a53cd9ee56e 100644 --- a/Mathlib/MeasureTheory/Constructions/Pi.lean +++ b/Mathlib/MeasureTheory/Constructions/Pi.lean @@ -90,7 +90,8 @@ theorem IsCountablySpanning.pi {C : ∀ i, Set (Set (α i))} (hC : ∀ i, IsCoun let e : ℕ → ι → ℕ := fun n => (@decode (ι → ℕ) _ n).iget refine ⟨fun n => Set.pi univ fun i => s i (e n i), fun n => mem_image_of_mem _ fun i _ => h1s i _, ?_⟩ - simp_rw [(surjective_decode_iget (ι → ℕ)).iUnion_comp fun x => Set.pi univ fun i => s i (x i), + simp_rw + [e, (surjective_decode_iget (ι → ℕ)).iUnion_comp fun x => Set.pi univ fun i => s i (x i), iUnion_univ_pi s, h2s, pi_univ] /-- The product of generated σ-algebras is the one generated by boxes, if both generating sets @@ -499,8 +500,8 @@ lemma pi_map_piOptionEquivProd {β : Option ι → Type*} [∀ i, MeasurableSpac simp only [mem_preimage, Set.mem_pi, mem_univ, forall_true_left, mem_prod] refine ⟨by tauto, fun _ i ↦ ?_⟩ rcases i <;> tauto - simp only [me.map_apply, univ_option, le_eq_subset, Finset.prod_insertNone, this, prod_prod, - pi_pi, mul_comm] + simp only [me.map_apply, univ_option, le_eq_subset, Finset.prod_insertNone, this, + prod_prod, pi_pi, mul_comm] section Intervals diff --git a/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean b/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean index d4071e4ee98fc..cdd41441ac246 100644 --- a/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean @@ -262,7 +262,7 @@ theorem AnalyticSet.iInter [hι : Nonempty ι] [Countable ι] [T2Space α] {s : choose x hx using A have xt : x ∈ t := by refine mem_iInter.2 fun n => ?_ - simp [hx] + simp [γ, t, F, hx] refine ⟨⟨x, xt⟩, ?_⟩ exact hx i₀ rw [← F_range] @@ -280,7 +280,7 @@ theorem AnalyticSet.iUnion [Countable ι] {s : ι → Set α} (hs : ∀ n, Analy let F : γ → α := fun ⟨n, x⟩ ↦ f n x have F_cont : Continuous F := continuous_sigma f_cont have F_range : range F = ⋃ n, s n := by - simp only [γ, range_sigma_eq_iUnion_range, f_range] + simp only [γ, F, range_sigma_eq_iUnion_range, f_range] rw [← F_range] exact analyticSet_range_of_polishSpace F_cont diff --git a/Mathlib/MeasureTheory/Covering/Besicovitch.lean b/Mathlib/MeasureTheory/Covering/Besicovitch.lean index 1e7a76eadda2d..8d6744400fbbc 100644 --- a/Mathlib/MeasureTheory/Covering/Besicovitch.lean +++ b/Mathlib/MeasureTheory/Covering/Besicovitch.lean @@ -566,7 +566,7 @@ theorem exist_finset_disjoint_balls_large_measure (μ : Measure α) [IsFiniteMea intro x hx obtain ⟨i, y, hxy, h'⟩ : ∃ (i : Fin N) (i_1 : ↥s), i_1 ∈ u i ∧ x ∈ ball (↑i_1) (r ↑i_1) := by - have : x ∈ range a.c := by simpa only [Subtype.range_coe_subtype, setOf_mem_eq] + have : x ∈ range a.c := by simpa only [a, Subtype.range_coe_subtype, setOf_mem_eq] simpa only [mem_iUnion, bex_def] using hu' this refine mem_iUnion.2 ⟨i, ⟨hx, ?_⟩⟩ simp only [v, exists_prop, mem_iUnion, SetCoe.exists, exists_and_right, Subtype.coe_mk] @@ -939,8 +939,8 @@ theorem exists_closedBall_covering_tsum_measure_le (μ : Measure α) [SFinite μ by_cases h'x : x ∈ s' · obtain ⟨i, y, ySi, xy⟩ : ∃ (i : Fin N) (y : ↥s'), y ∈ S i ∧ x ∈ ball (y : α) (r1 y) := by have A : x ∈ range q.c := by - simpa only [not_exists, exists_prop, mem_iUnion, mem_closedBall, not_and, not_le, - mem_setOf_eq, Subtype.range_coe_subtype, mem_diff] using h'x + simpa only [q, not_exists, exists_prop, mem_iUnion, mem_closedBall, not_and, + not_le, mem_setOf_eq, Subtype.range_coe_subtype, mem_diff] using h'x simpa only [mem_iUnion, mem_image, bex_def] using hS A refine mem_iUnion₂.2 ⟨y, Or.inr ?_, ?_⟩ · simp only [mem_iUnion, mem_image] diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Unique.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Unique.lean index aa7f322d0889d..3896745fc04fb 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Unique.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Unique.lean @@ -77,7 +77,7 @@ theorem Lp.ae_eq_zero_of_forall_setIntegral_eq_zero' (hm : m ≤ m0) (f : Lp E' (hf_meas : AEStronglyMeasurable' m f μ) : f =ᵐ[μ] 0 := by let f_meas : lpMeas E' 𝕜 m p μ := ⟨f, hf_meas⟩ -- Porting note: `simp only` does not call `rfl` to try to close the goal. See https://github.com/leanprover-community/mathlib4/issues/5025 - have hf_f_meas : f =ᵐ[μ] f_meas := by simp only [Subtype.coe_mk]; rfl + have hf_f_meas : f =ᵐ[μ] f_meas := by simp only [f_meas, Subtype.coe_mk]; rfl refine hf_f_meas.trans ?_ refine lpMeas.ae_eq_zero_of_forall_setIntegral_eq_zero hm f_meas hp_ne_zero hp_ne_top ?_ ?_ · intro s hs hμs diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index 1d8f510752db3..7d6a28d6d78b1 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -1479,7 +1479,7 @@ private theorem eLpNorm'_sum_norm_sub_le_tsum_of_cauchy_eLpNorm' {f : ℕ → α ∀ n, (fun x => ∑ i ∈ Finset.range (n + 1), ‖f (i + 1) x - f i x‖) = ∑ i ∈ Finset.range (n + 1), f_norm_diff i := - fun n => funext fun x => by simp + fun n => funext fun x => by simp [f_norm_diff] rw [hgf_norm_diff] refine (eLpNorm'_sum_le (fun i _ => ((hf (i + 1)).sub (hf i)).norm) hp1).trans ?_ simp_rw [eLpNorm'_norm] diff --git a/Mathlib/MeasureTheory/Group/GeometryOfNumbers.lean b/Mathlib/MeasureTheory/Group/GeometryOfNumbers.lean index 4f7644744e85c..72cc5e33e2df2 100644 --- a/Mathlib/MeasureTheory/Group/GeometryOfNumbers.lean +++ b/Mathlib/MeasureTheory/Group/GeometryOfNumbers.lean @@ -135,6 +135,6 @@ theorem exists_ne_zero_mem_lattice_of_measure_mul_two_pow_le_measure [NormedAddC 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) - simp [(exists_seq_strictAnti_tendsto (0 : ℝ≥0)).choose_spec.2.1 n] + simp [u, K, S, Z, (exists_seq_strictAnti_tendsto (0 : ℝ≥0)).choose_spec.2.1 n] end MeasureTheory diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index 00aeccd87421d..65a635f05396d 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -1168,7 +1168,8 @@ lemma integral_tendsto_of_tendsto_of_monotone {μ : Measure α} {f : ℕ → α simp [f', ha (zero_le n)] have hf'_meas : ∀ n, Integrable (f' n) μ := fun n ↦ (hf n).sub (hf 0) suffices Tendsto (fun n ↦ ∫ x, f' n x ∂μ) atTop (𝓝 (∫ x, (F - f 0) x ∂μ)) by - simp_rw [integral_sub (hf _) (hf _), integral_sub' hF (hf 0), tendsto_sub_const_iff] at this + simp_rw [f', integral_sub (hf _) (hf _), integral_sub' hF (hf 0), + tendsto_sub_const_iff] at this exact this have hF_ge : 0 ≤ᵐ[μ] fun x ↦ (F - f 0) x := by filter_upwards [h_tendsto, h_mono] with x hx_tendsto hx_mono diff --git a/Mathlib/MeasureTheory/Integral/Marginal.lean b/Mathlib/MeasureTheory/Integral/Marginal.lean index 4e19390aa37fb..2c027ef3bc5e7 100644 --- a/Mathlib/MeasureTheory/Integral/Marginal.lean +++ b/Mathlib/MeasureTheory/Integral/Marginal.lean @@ -122,8 +122,9 @@ theorem lmarginal_singleton (f : (∀ i, π i) → ℝ≥0∞) (i : δ) : ext1 x calc (∫⋯∫⁻_{i}, f ∂μ) x = ∫⁻ (y : π (default : α)), f (updateFinset x {i} (e y)) ∂μ (default : α) := by - simp_rw [lmarginal, measurePreserving_piUnique (fun j : ({i} : Finset δ) ↦ μ j) |>.symm _ - |>.lintegral_map_equiv] + simp_rw [lmarginal, + measurePreserving_piUnique (fun j : ({i} : Finset δ) ↦ μ j) |>.symm _ + |>.lintegral_map_equiv] _ = ∫⁻ xᵢ, f (Function.update x i xᵢ) ∂μ i := by simp [update_eq_updateFinset]; rfl variable {μ} in diff --git a/Mathlib/MeasureTheory/Integral/MeanInequalities.lean b/Mathlib/MeasureTheory/Integral/MeanInequalities.lean index 424952fb5cb92..1cfbb6e4f677d 100644 --- a/Mathlib/MeasureTheory/Integral/MeanInequalities.lean +++ b/Mathlib/MeasureTheory/Integral/MeanInequalities.lean @@ -221,7 +221,7 @@ theorem lintegral_prod_norm_pow_le {α ι : Type*} [MeasurableSpace α] {μ : Me calc ∫⁻ a, ∏ i ∈ insert i₀ s, f i a ^ p i ∂μ = ∫⁻ a, f i₀ a ^ p i₀ * ∏ i ∈ s, f i a ^ p i ∂μ := by simp [hi₀] _ = ∫⁻ a, f i₀ a ^ p i₀ * (∏ i ∈ s, f i a ^ q i) ^ (1 - p i₀) ∂μ := by - simp [← ENNReal.prod_rpow_of_nonneg hpi₀, ← ENNReal.rpow_mul, + simp [q, ← ENNReal.prod_rpow_of_nonneg hpi₀, ← ENNReal.rpow_mul, div_mul_cancel₀ (h := h2pi₀)] _ ≤ (∫⁻ a, f i₀ a ∂μ) ^ p i₀ * (∫⁻ a, ∏ i ∈ s, f i a ^ q i ∂μ) ^ (1 - p i₀) := by apply ENNReal.lintegral_mul_norm_pow_le @@ -234,7 +234,7 @@ theorem lintegral_prod_norm_pow_le {α ι : Type*} [MeasurableSpace α] {μ : Me gcongr -- behavior of gcongr is heartbeat-dependent, which makes code really fragile... exact ih (fun i hi ↦ hf i <| mem_insert_of_mem hi) hq h2q _ = (∫⁻ a, f i₀ a ∂μ) ^ p i₀ * ∏ i ∈ s, (∫⁻ a, f i a ∂μ) ^ p i := by - simp [← ENNReal.prod_rpow_of_nonneg hpi₀, ← ENNReal.rpow_mul, + simp [q, ← ENNReal.prod_rpow_of_nonneg hpi₀, ← ENNReal.rpow_mul, div_mul_cancel₀ (h := h2pi₀)] _ = ∏ i ∈ insert i₀ s, (∫⁻ a, f i a ∂μ) ^ p i := by simp [hi₀] diff --git a/Mathlib/MeasureTheory/Integral/SetToL1.lean b/Mathlib/MeasureTheory/Integral/SetToL1.lean index 2a0d2d70bd546..7c6e7f8e1b323 100644 --- a/Mathlib/MeasureTheory/Integral/SetToL1.lean +++ b/Mathlib/MeasureTheory/Integral/SetToL1.lean @@ -1365,7 +1365,7 @@ theorem tendsto_setToFun_of_L1 (hT : DominatedFinMeasAdditive μ T C) {ι} (f : filter_upwards [hfsi] with i hi refine lintegral_congr_ae ?_ filter_upwards [hi.coeFn_toL1, hfi.coeFn_toL1] with x hxi hxf - simp_rw [F_lp, dif_pos hi, hxi, hxf] + simp_rw [F_lp, dif_pos hi, hxi, f_lp, hxf] suffices Tendsto (fun i => setToFun μ T hT (F_lp i)) l (𝓝 (setToFun μ T hT f)) by refine (tendsto_congr' ?_).mp this filter_upwards [hfsi] with i hi diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean b/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean index 0a63d44706834..12eef5dd443c3 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean @@ -722,7 +722,7 @@ noncomputable def schroederBernstein {f : α → β} {g : β → α} (hf : Measu have : Aᶜ = g '' Bᶜ := by apply compl_injective rw [← Afp] - simp + simp [F, B] rw [this] exact (hg.equivImage _).symm have Fmono : ∀ {A B}, A ⊆ B → F A ⊆ F B := fun h => diff --git a/Mathlib/MeasureTheory/Measure/AEMeasurable.lean b/Mathlib/MeasureTheory/Measure/AEMeasurable.lean index c159b4c1db9d1..655b839adfe70 100644 --- a/Mathlib/MeasureTheory/Measure/AEMeasurable.lean +++ b/Mathlib/MeasureTheory/Measure/AEMeasurable.lean @@ -194,7 +194,7 @@ theorem exists_ae_eq_range_subset (H : AEMeasurable f μ) {t : Set β} (ht : ∀ exact H.ae_eq_mk.and ht filter_upwards [compl_mem_ae_iff.2 A] with x hx rw [mem_compl_iff] at hx - simp only [g, hx, piecewise_eq_of_not_mem, not_false_iff] + simp only [s, g, hx, piecewise_eq_of_not_mem, not_false_iff] contrapose! hx apply subset_toMeasurable simp only [hx, mem_compl_iff, mem_setOf_eq, false_and, not_false_iff] diff --git a/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean b/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean index 9f84eee98f624..c9da154064847 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean @@ -105,14 +105,14 @@ theorem parallelepiped_orthonormalBasis_one_dim (b : OrthonormalBasis ι ℝ ℝ · intro x hx refine ⟨x 0, ⟨hx.1 0, hx.2 0⟩, ?_⟩ ext j - simp only [Subsingleton.elim j 0] + simp only [F, Subsingleton.elim j 0] · rintro x ⟨y, hy, rfl⟩ exact ⟨fun _j => hy.1, fun _j => hy.2⟩ rcases orthonormalBasis_one_dim (b.reindex e) with (H | H) · left simp_rw [parallelepiped, H, A, Algebra.id.smul_eq_mul, mul_one] - simp only [Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton, ← image_comp, - Function.comp_apply, image_id'] + simp only [F, Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton, + ← image_comp, Function.comp_apply, image_id'] · right simp_rw [H, parallelepiped, Algebra.id.smul_eq_mul, A] simp only [F, Finset.univ_unique, Fin.default_eq_zero, mul_neg, mul_one, Finset.sum_neg_distrib, diff --git a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean index 5d1a0d0477fb8..377e30300f09c 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean @@ -169,7 +169,7 @@ lemma integral_isMulLeftInvariant_isMulRightInvariant_combo calc ∫ x, f x ∂μ = ∫ x, f x * (D x)⁻¹ * D x ∂μ := by congr with x; rw [mul_assoc, inv_mul_cancel₀ (D_pos x).ne', mul_one] - _ = ∫ x, (∫ y, f x * (D x)⁻¹ * g (y⁻¹ * x) ∂ν) ∂μ := by simp_rw [integral_mul_left] + _ = ∫ x, (∫ y, f x * (D x)⁻¹ * g (y⁻¹ * x) ∂ν) ∂μ := by simp_rw [D, integral_mul_left] _ = ∫ y, (∫ x, f x * (D x)⁻¹ * g (y⁻¹ * x) ∂μ) ∂ν := by apply integral_integral_swap_of_hasCompactSupport · apply Continuous.mul diff --git a/Mathlib/MeasureTheory/Measure/Prod.lean b/Mathlib/MeasureTheory/Measure/Prod.lean index 4a67f092f79cf..b1b1cf0cd1ea5 100644 --- a/Mathlib/MeasureTheory/Measure/Prod.lean +++ b/Mathlib/MeasureTheory/Measure/Prod.lean @@ -194,7 +194,7 @@ theorem prod_prod (s : Set α) (t : Set β) : μ.prod ν (s ×ˢ t) = μ s * ν μ.prod ν (s ×ˢ t) ≤ μ.prod ν (S ×ˢ T) := by gcongr <;> apply subset_toMeasurable _ = μ S * ν T := by rw [prod_apply hSTm] - simp_rw [mk_preimage_prod_right_eq_if, measure_if, + simp_rw [S, mk_preimage_prod_right_eq_if, measure_if, lintegral_indicator (measurableSet_toMeasurable _ _), lintegral_const, restrict_apply_univ, mul_comm] _ = μ s * ν t := by rw [measure_toMeasurable, measure_toMeasurable] diff --git a/Mathlib/MeasureTheory/Measure/WithDensity.lean b/Mathlib/MeasureTheory/Measure/WithDensity.lean index f8c9285e8bbbf..f2b8b222f5c26 100644 --- a/Mathlib/MeasureTheory/Measure/WithDensity.lean +++ b/Mathlib/MeasureTheory/Measure/WithDensity.lean @@ -385,7 +385,7 @@ theorem lintegral_withDensity_eq_lintegral_mul₀' {μ : Measure α} {f : α → rw [EventuallyEq, ae_withDensity_iff_ae_restrict hf.measurable_mk] at Z filter_upwards [Z] intro x hx - simp only [hx, Pi.mul_apply] + simp only [g', hx, Pi.mul_apply] · have M : MeasurableSet { x : α | f' x ≠ 0 }ᶜ := (hf.measurable_mk (measurableSet_singleton 0).compl).compl filter_upwards [ae_restrict_mem M] @@ -396,7 +396,7 @@ theorem lintegral_withDensity_eq_lintegral_mul₀' {μ : Measure α} {f : α → apply lintegral_congr_ae filter_upwards [hf.ae_eq_mk] intro x hx - simp only [hx, Pi.mul_apply] + simp only [f', hx, Pi.mul_apply] lemma setLIntegral_withDensity_eq_lintegral_mul₀' {μ : Measure α} {f : α → ℝ≥0∞} (hf : AEMeasurable f μ) {g : α → ℝ≥0∞} (hg : AEMeasurable g (μ.withDensity f)) @@ -483,7 +483,7 @@ theorem lintegral_withDensity_eq_lintegral_mul_non_measurable₀ (μ : Measure apply lintegral_congr_ae filter_upwards [hf.ae_eq_mk] intro x hx - simp only [hx, Pi.mul_apply] + simp only [f', hx, Pi.mul_apply] theorem setLIntegral_withDensity_eq_setLIntegral_mul_non_measurable₀ (μ : Measure α) {f : α → ℝ≥0∞} {s : Set α} (hf : AEMeasurable f (μ.restrict s)) (g : α → ℝ≥0∞) diff --git a/Mathlib/MeasureTheory/OuterMeasure/OfFunction.lean b/Mathlib/MeasureTheory/OuterMeasure/OfFunction.lean index ac7e4d2a1aaca..116867bde48c2 100644 --- a/Mathlib/MeasureTheory/OuterMeasure/OfFunction.lean +++ b/Mathlib/MeasureTheory/OuterMeasure/OfFunction.lean @@ -124,7 +124,7 @@ theorem ofFunction_le (s : Set α) : OuterMeasure.ofFunction m m_empty s ≤ m s le_of_eq <| tsum_eq_single 0 <| by rintro (_ | i) · simp - · simp [m_empty] + · simp [f, m_empty] theorem ofFunction_eq (s : Set α) (m_mono : ∀ ⦃t : Set α⦄, s ⊆ t → m s ≤ m t) (m_subadd : ∀ s : ℕ → Set α, m (⋃ i, s i) ≤ ∑' i, m (s i)) : diff --git a/Mathlib/ModelTheory/Fraisse.lean b/Mathlib/ModelTheory/Fraisse.lean index 11b8204c05fa6..f782029a8eaa0 100644 --- a/Mathlib/ModelTheory/Fraisse.lean +++ b/Mathlib/ModelTheory/Fraisse.lean @@ -389,7 +389,7 @@ protected theorem isExtensionPair : L.IsExtensionPair M N := by refine ⟨⟨⟨S, g.toHom.range, g.equivRange⟩, S_FG⟩, subset_closure.trans (le_sup_right : _ ≤ S) (mem_singleton m), ⟨le_sup_left, ?_⟩⟩ ext - simp [Subtype.mk_le_mk, PartialEquiv.le_def, g_eq] + simp [S, Subtype.mk_le_mk, PartialEquiv.le_def, g_eq] /-- The Fraïssé limit of a class is unique, in that any two Fraïssé limits are isomorphic. -/ theorem nonempty_equiv : Nonempty (M ≃[L] N) := by diff --git a/Mathlib/NumberTheory/LSeries/HurwitzZetaOdd.lean b/Mathlib/NumberTheory/LSeries/HurwitzZetaOdd.lean index 7d60dc787c117..f59d387a00ae8 100644 --- a/Mathlib/NumberTheory/LSeries/HurwitzZetaOdd.lean +++ b/Mathlib/NumberTheory/LSeries/HurwitzZetaOdd.lean @@ -453,7 +453,7 @@ lemma hasSum_int_completedHurwitzZetaOdd (a : ℝ) {s : ℂ} (hs : 1 < re s) : rwa [summable_one_div_int_add_rpow] have := mellin_div_const .. ▸ hasSum_mellin_pi_mul_sq' (zero_lt_one.trans hs) hF h_sum refine this.congr_fun fun n ↦ ?_ - simp only [c, mul_one_div, div_mul_eq_mul_div, div_right_comm] + simp only [r, c, mul_one_div, div_mul_eq_mul_div, div_right_comm] /-! ## Non-completed zeta functions diff --git a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean index b25bd5ca4f8ea..360acae7e6d20 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean @@ -246,7 +246,7 @@ theorem quadraticChar_card_sqrts (hF : ringChar F ≠ 2) (a : F) : ext1 -- Porting note(https://github.com/leanprover-community/mathlib4/issues/5026): -- added (Set.mem_toFinset), Set.mem_setOf - simp only [(Set.mem_toFinset), Set.mem_setOf, not_mem_empty, iff_false] + simp only [s, (Set.mem_toFinset), Set.mem_setOf, not_mem_empty, iff_false] rw [isSquare_iff_exists_sq] at h exact fun h' ↦ h ⟨_, h'.symm⟩ diff --git a/Mathlib/NumberTheory/MaricaSchoenheim.lean b/Mathlib/NumberTheory/MaricaSchoenheim.lean index a88458695d67d..d1dbc047e3b2a 100644 --- a/Mathlib/NumberTheory/MaricaSchoenheim.lean +++ b/Mathlib/NumberTheory/MaricaSchoenheim.lean @@ -47,13 +47,13 @@ lemma grahamConjecture_of_squarefree {n : ℕ} (f : ℕ → ℕ) (hf' : ∀ k < _ < n := tsub_lt_self hn.bot_lt zero_lt_one · rw [Finset.card_image_of_injOn, card_Iio] simpa using prod_primeFactors_invOn_squarefree.2.injOn.comp hf.injOn hf' - · simp only [forall_mem_diffs, forall_image, mem_Ioo, mem_Iio] + · simp only [𝒜, forall_mem_diffs, forall_image, mem_Ioo, mem_Iio] rintro i hi j hj rw [← primeFactors_div_gcd (hf' _ hi) (hf' _ hj).ne_zero, prod_primeFactors_of_squarefree <| hf'' _ hi _] exact ⟨Nat.div_pos (gcd_le_left _ (hf' _ hi).ne_zero.bot_lt) <| Nat.gcd_pos_of_pos_left _ (hf' _ hi).ne_zero.bot_lt, Nat.div_lt_of_lt_mul <| this _ hi _ hj⟩ - · simp only [Set.InjOn, mem_coe, forall_mem_diffs, forall_image, mem_Ioo, mem_Iio] + · simp only [𝒜, Set.InjOn, mem_coe, forall_mem_diffs, forall_image, mem_Ioo, mem_Iio] rintro a ha b hb c hc d hd rw [← primeFactors_div_gcd (hf' _ ha) (hf' _ hb).ne_zero, ← primeFactors_div_gcd (hf' _ hc) (hf' _ hd).ne_zero, prod_primeFactors_of_squarefree (hf'' _ ha _), diff --git a/Mathlib/NumberTheory/Multiplicity.lean b/Mathlib/NumberTheory/Multiplicity.lean index ab751110fdab3..ec4fd6442cf28 100644 --- a/Mathlib/NumberTheory/Multiplicity.lean +++ b/Mathlib/NumberTheory/Multiplicity.lean @@ -90,7 +90,7 @@ theorem odd_sq_dvd_geom_sum₂_sub (hp : Odd p) : (Ideal.Quotient.mk (span {s})) (∑ i ∈ range p, (a + (p : R) * b) ^ i * a ^ (p - 1 - i)) = ∑ i ∈ Finset.range p, mk (span {s}) ((a ^ (i - 1) * (↑p * b) * ↑i + a ^ i) * a ^ (p - 1 - i)) := by - simp_rw [RingHom.map_geom_sum₂, ← map_pow, h1, ← _root_.map_mul] + simp_rw [s, RingHom.map_geom_sum₂, ← map_pow, h1, ← _root_.map_mul] _ = mk (span {s}) (∑ x ∈ Finset.range p, a ^ (x - 1) * (a ^ (p - 1 - x) * (↑p * (b * ↑x)))) + @@ -139,7 +139,7 @@ theorem odd_sq_dvd_geom_sum₂_sub (hp : Odd p) : rw [mul_assoc, mul_assoc] refine mul_eq_zero_of_left ?_ _ refine Ideal.Quotient.eq_zero_iff_mem.mpr ?_ - simp [mem_span_singleton] + simp [s, mem_span_singleton] section IntegralDomain diff --git a/Mathlib/NumberTheory/NumberField/Embeddings.lean b/Mathlib/NumberTheory/NumberField/Embeddings.lean index 4b72199ed61ab..9cda9eb7b84d8 100644 --- a/Mathlib/NumberTheory/NumberField/Embeddings.lean +++ b/Mathlib/NumberTheory/NumberField/Embeddings.lean @@ -638,7 +638,7 @@ lemma comap_surjective [Algebra k K] [Algebra.IsAlgebraic k K] : Function.Surjective (comap · (algebraMap k K)) := fun w ↦ letI := w.embedding.toAlgebra ⟨mk (IsAlgClosed.lift (M := ℂ) (R := k)).toRingHom, - by simp [comap_mk, RingHom.algebraMap_toAlgebra]⟩ + by simp [this, comap_mk, RingHom.algebraMap_toAlgebra]⟩ lemma mult_comap_le (f : k →+* K) (w : InfinitePlace K) : mult (w.comap f) ≤ mult w := by rw [mult, mult] diff --git a/Mathlib/NumberTheory/Padics/Hensel.lean b/Mathlib/NumberTheory/Padics/Hensel.lean index 74e302eb5dbe2..cf817d16f0a7c 100644 --- a/Mathlib/NumberTheory/Padics/Hensel.lean +++ b/Mathlib/NumberTheory/Padics/Hensel.lean @@ -432,7 +432,7 @@ private theorem soln_unique (z : ℤ_[p]) (hev : F.eval z = 0) _ ≤ 1 * ‖h‖ := by rw [PadicInt.norm_mul] exact mul_le_mul_of_nonneg_right (PadicInt.norm_le_one _) (norm_nonneg _) - _ = ‖z - soln‖ := by simp + _ = ‖z - soln‖ := by simp [h] _ < ‖F.derivative.eval soln‖ := by rw [soln_deriv_norm]; apply soln_dist ) eq_of_sub_eq_zero (by rw [← this]) diff --git a/Mathlib/NumberTheory/Padics/PadicIntegers.lean b/Mathlib/NumberTheory/Padics/PadicIntegers.lean index db88a79d9130f..33bdb2db15211 100644 --- a/Mathlib/NumberTheory/Padics/PadicIntegers.lean +++ b/Mathlib/NumberTheory/Padics/PadicIntegers.lean @@ -615,8 +615,8 @@ instance isFractionRing : IsFractionRing ℤ_[p] ℚ_[p] where use (⟨a, le_of_eq ha_norm⟩, ⟨(p ^ n : ℤ_[p]), mem_nonZeroDivisors_iff_ne_zero.mpr (NeZero.ne _)⟩) - simp only [map_pow, map_natCast, algebraMap_apply, PadicInt.coe_pow, PadicInt.coe_natCast, - Subtype.coe_mk, Nat.cast_pow] + simp only [a, map_pow, map_natCast, algebraMap_apply, PadicInt.coe_pow, + PadicInt.coe_natCast, Subtype.coe_mk, Nat.cast_pow] exists_of_eq := by simp_rw [algebraMap_apply, Subtype.coe_inj] exact fun h => ⟨1, by rw [h]⟩ diff --git a/Mathlib/NumberTheory/Padics/RingHoms.lean b/Mathlib/NumberTheory/Padics/RingHoms.lean index 2fed1c971dbfe..53244367555c2 100644 --- a/Mathlib/NumberTheory/Padics/RingHoms.lean +++ b/Mathlib/NumberTheory/Padics/RingHoms.lean @@ -119,7 +119,8 @@ theorem norm_sub_modPart (h : ‖(r : ℚ_[p])‖ ≤ 1) : ‖(⟨r, h⟩ - modP rw [norm_lt_one_iff_dvd, ← (isUnit_den r h).dvd_mul_right] suffices ↑p ∣ r.num - n * r.den by convert (Int.castRingHom ℤ_[p]).map_dvd this - simp only [sub_mul, Int.cast_natCast, eq_intCast, Int.cast_mul, sub_left_inj, Int.cast_sub] + simp only [n, sub_mul, Int.cast_natCast, eq_intCast, Int.cast_mul, sub_left_inj, + Int.cast_sub] apply Subtype.coe_injective simp only [coe_mul, Subtype.coe_mk, coe_natCast] rw_mod_cast [@Rat.mul_den_eq_num r] diff --git a/Mathlib/NumberTheory/WellApproximable.lean b/Mathlib/NumberTheory/WellApproximable.lean index 7b72f49ff60e8..a4baa65fee7dd 100644 --- a/Mathlib/NumberTheory/WellApproximable.lean +++ b/Mathlib/NumberTheory/WellApproximable.lean @@ -274,7 +274,7 @@ theorem addWellApproximable_ae_empty_or_univ (δ : ℕ → ℝ) (hδ : Tendsto exact fun contra => h_ndiv (mul_dvd_mul_left p contra) replace h_div : n / p * p = n := Nat.div_mul_cancel h_div have hf : f = (fun y => x + y) ∘ fun y => p • y := by - ext; simp [add_comm x] + ext; simp [f, add_comm x] simp_rw [Function.comp_apply, le_eq_subset] rw [sSupHom.setImage_toFun, hf, image_comp] have := @monotone_image 𝕊 𝕊 fun y => x + y @@ -309,7 +309,8 @@ theorem addWellApproximable_ae_empty_or_univ (δ : ℕ → ℝ) (hδ : Tendsto · cases' hA p with _ h; · contradiction simp only [h, union_ae_eq_univ_of_ae_eq_univ_left] · cases' hB p with _ h; · contradiction - simp only [h, union_ae_eq_univ_of_ae_eq_univ_left, union_ae_eq_univ_of_ae_eq_univ_right] + simp only [h, union_ae_eq_univ_of_ae_eq_univ_left, + union_ae_eq_univ_of_ae_eq_univ_right] /-- A general version of **Dirichlet's approximation theorem**. @@ -335,7 +336,7 @@ lemma _root_.NormedAddCommGroup.exists_norm_nsmul_le {A : Type*} rw [← (isClosed_iUnion_of_finite hB).measure_eq_univ_iff_eq (μ := μ)] refine le_antisymm (μ.mono (subset_univ _)) ?_ simp_rw [measure_iUnion h (fun _ ↦ measurableSet_closedBall), tsum_fintype, - μ.addHaar_closedBall_center, Finset.sum_const, Finset.card_univ, Nat.card_fintypeIcc, + B, μ.addHaar_closedBall_center, Finset.sum_const, Finset.card_univ, Nat.card_fintypeIcc, tsub_zero] exact hδ replace hδ : 0 ≤ δ/2 := by diff --git a/Mathlib/Order/Category/NonemptyFinLinOrd.lean b/Mathlib/Order/Category/NonemptyFinLinOrd.lean index 1bf3149bd7abe..ea21848f75410 100644 --- a/Mathlib/Order/Category/NonemptyFinLinOrd.lean +++ b/Mathlib/Order/Category/NonemptyFinLinOrd.lean @@ -178,7 +178,7 @@ theorem epi_iff_surjective {A B : NonemptyFinLinOrd.{u}} (f : A ⟶ B) : exact h₂ (le_of_lt h₁) · exfalso exact hm a (eq_of_le_of_not_lt h₂ h₁) - simp [Y, DFunLike.coe] at h + simp [Y, p₁, p₂, DFunLike.coe] at h · intro h exact ConcreteCategory.epi_of_surjective f h diff --git a/Mathlib/Order/KrullDimension.lean b/Mathlib/Order/KrullDimension.lean index 1b435cca65877..a6ba1d3f81334 100644 --- a/Mathlib/Order/KrullDimension.lean +++ b/Mathlib/Order/KrullDimension.lean @@ -750,7 +750,7 @@ lemma krullDim_int : krullDim ℤ = ⊤ := krullDim_of_noMaxOrder .. exact compare_gt_iff_gt.mp rfl) step := fun i => by simpa [WithBot.unbot_lt_iff] using p.step ⟨i + 1, by omega⟩ } have hlast' : p'.last = x := by - simp only [RelSeries.last, Fin.val_last, WithBot.unbot_eq_iff, ← hlast, Fin.last] + simp only [p', RelSeries.last, Fin.val_last, WithBot.unbot_eq_iff, ← hlast, Fin.last] congr omega suffices p'.length ≤ height p'.last by @@ -780,7 +780,7 @@ lemma krullDim_int : krullDim ℤ = ⊤ := krullDim_of_noMaxOrder .. simp [hlast]) step := fun i => by simpa only [WithTop.untop_lt_iff, WithTop.coe_untop] using p.step i } have hlast' : p'.last = x := by - simp only [RelSeries.last, Fin.val_last, WithTop.untop_eq_iff, ← hlast] + simp only [p', RelSeries.last, Fin.val_last, WithTop.untop_eq_iff, ← hlast] suffices p'.length ≤ height p'.last by rw [hlast'] at this simpa [p'] using this diff --git a/Mathlib/Probability/StrongLaw.lean b/Mathlib/Probability/StrongLaw.lean index 4740c2314e1fb..9094e67f68a42 100644 --- a/Mathlib/Probability/StrongLaw.lean +++ b/Mathlib/Probability/StrongLaw.lean @@ -436,7 +436,7 @@ theorem strong_law_aux1 {c : ℝ} (c_one : 1 < c) {ε : ℝ} (εpos : 0 < ε) : apply mul_le_mul_of_nonneg_right _ (variance_nonneg _ _) convert sum_div_nat_floor_pow_sq_le_div_sq N (Nat.cast_pos.2 hj) c_one using 2 · simp only [Nat.cast_lt] - · simp only [one_div] + · simp only [Y, S, u, C, one_div] _ = c ^ 5 * (c - 1)⁻¹ ^ 3 * ∑ j ∈ range (u (N - 1)), ((j : ℝ) ^ 2)⁻¹ * Var[Y j] := by simp_rw [mul_sum, div_eq_mul_inv, mul_assoc] _ ≤ c ^ 5 * (c - 1)⁻¹ ^ 3 * (2 * 𝔼[X 0]) := by @@ -475,7 +475,7 @@ theorem strong_law_aux1 {c : ℝ} (c_one : 1 < c) {ε : ℝ} (εpos : 0 < ε) : ENNReal.ofReal_lt_top filter_upwards [ae_eventually_not_mem I4.ne] with ω hω simp_rw [S, not_le, mul_comm, sum_apply] at hω - convert hω; simp only [sum_apply] + convert hω; simp only [Y, S, u, C, sum_apply] include hint hindep hident hnonneg in /- The truncation of `Xᵢ` up to `i` satisfies the strong law of large numbers @@ -634,8 +634,8 @@ theorem strong_law_ae_real {Ω : Type*} {m : MeasurableSpace Ω} {μ : Measure convert hωpos.sub hωneg using 2 · simp only [pos, neg, ← sub_div, ← sum_sub_distrib, max_zero_sub_max_neg_zero_eq_self, Function.comp_apply] - · simp only [← integral_sub hint.pos_part hint.neg_part, max_zero_sub_max_neg_zero_eq_self, - Function.comp_apply, mΩ] + · simp only [← integral_sub hint.pos_part hint.neg_part, + max_zero_sub_max_neg_zero_eq_self, Function.comp_apply, mΩ] end StrongLawAeReal @@ -684,7 +684,7 @@ lemma strong_law_ae_simpleFunc_comp (X : ℕ → Ω → E) (h' : Measurable (X 0 simp simp only [I, integral_smul_const] convert Tendsto.smul_const hω c using 1 - simp [Y, ← sum_smul, smul_smul] + simp [F, Y, ← sum_smul, smul_smul] · rintro φ ψ - hφ hψ filter_upwards [hφ, hψ] with ω hωφ hωψ convert hωφ.add hωψ using 1 diff --git a/Mathlib/RingTheory/Adjoin/PowerBasis.lean b/Mathlib/RingTheory/Adjoin/PowerBasis.lean index 2a3bbcac0b345..5b1c9fd4bce17 100644 --- a/Mathlib/RingTheory/Adjoin/PowerBasis.lean +++ b/Mathlib/RingTheory/Adjoin/PowerBasis.lean @@ -94,7 +94,7 @@ theorem repr_gen_pow_isIntegral (hB : IsIntegral R B.gen) [IsDomain S] let Q := X ^ n %ₘ minpoly R B.gen have : B.gen ^ n = aeval B.gen Q := by rw [← @aeval_X_pow R _ _ _ _ B.gen, ← modByMonic_add_div (X ^ n) (minpoly.monic hB)] - simp + simp [Q] by_cases hQ : Q = 0 · simp [this, hQ, isIntegral_zero] have hlt : Q.natDegree < B.dim := by diff --git a/Mathlib/RingTheory/AlgebraicIndependent/RankAndCardinality.lean b/Mathlib/RingTheory/AlgebraicIndependent/RankAndCardinality.lean index da5e72bcfd342..1c736c384c401 100644 --- a/Mathlib/RingTheory/AlgebraicIndependent/RankAndCardinality.lean +++ b/Mathlib/RingTheory/AlgebraicIndependent/RankAndCardinality.lean @@ -44,7 +44,7 @@ theorem IsTranscendenceBasis.lift_cardinalMk_eq_max_lift {ι : Type w} {x : ι → E} [Nonempty ι] (hx : IsTranscendenceBasis F x) : lift.{max u w} #E = lift.{max v w} #F ⊔ lift.{max u v} #ι ⊔ ℵ₀ := by let K := Algebra.adjoin F (Set.range x) - suffices #E = #K by simp [this, ← lift_mk_eq'.2 ⟨hx.1.aevalEquiv.toEquiv⟩] + suffices #E = #K by simp [K, this, ← lift_mk_eq'.2 ⟨hx.1.aevalEquiv.toEquiv⟩] haveI : Algebra.IsAlgebraic K E := hx.isAlgebraic refine le_antisymm ?_ (mk_le_of_injective Subtype.val_injective) haveI : Infinite K := hx.1.aevalEquiv.infinite_iff.1 inferInstance @@ -60,7 +60,7 @@ theorem IsTranscendenceBasis.lift_rank_eq_max_lift MvRatFunc.rank_eq_max_lift, lift_max, lift_max, lift_lift, lift_lift, lift_aleph0] refine mul_eq_left le_sup_right ((lift_le.2 ((rank_le_card K E).trans (Algebra.IsAlgebraic.cardinalMk_le_max K E))).trans_eq ?_) (by simp [rank_pos.ne']) - simp [← lift_mk_eq'.2 ⟨hx.1.aevalEquivField.toEquiv⟩] + simp [K, ← lift_mk_eq'.2 ⟨hx.1.aevalEquivField.toEquiv⟩] theorem Algebra.Transcendental.rank_eq_cardinalMk (F : Type u) (E : Type v) [Field F] [Field E] [Algebra F E] [Algebra.Transcendental F E] : diff --git a/Mathlib/RingTheory/DedekindDomain/PID.lean b/Mathlib/RingTheory/DedekindDomain/PID.lean index 69bf559e5c7f8..673e136d6943e 100644 --- a/Mathlib/RingTheory/DedekindDomain/PID.lean +++ b/Mathlib/RingTheory/DedekindDomain/PID.lean @@ -110,7 +110,7 @@ theorem FractionalIdeal.isPrincipal.of_finite_maximals_of_inv {A : Type*} [CommR let s := hf.toFinset haveI := Classical.decEq (Ideal R) have coprime : ∀ M ∈ s, ∀ M' ∈ s.erase M, M ⊔ M' = ⊤ := by - simp_rw [Finset.mem_erase, hf.mem_toFinset] + simp_rw [s, Finset.mem_erase, hf.mem_toFinset] rintro M hM M' ⟨hne, hM'⟩ exact Ideal.IsMaximal.coprime_of_ne hM hM' hne.symm have nle : ∀ M ∈ s, ¬⨅ M' ∈ s.erase M, M' ≤ M := fun M hM => diff --git a/Mathlib/RingTheory/IsTensorProduct.lean b/Mathlib/RingTheory/IsTensorProduct.lean index bbe4d72567496..1023fe388ff81 100644 --- a/Mathlib/RingTheory/IsTensorProduct.lean +++ b/Mathlib/RingTheory/IsTensorProduct.lean @@ -334,7 +334,7 @@ lemma IsBaseChange.of_comp {f : M →ₗ[R] N} (hf : IsBaseChange S f) {h : N let q : O →ₗ[T] Q := hc.lift r' refine ⟨q, ?_, ?_⟩ · apply hf.algHom_ext' - simp [LinearMap.comp_assoc, hc.lift_comp] + simp [r', q, LinearMap.comp_assoc, hc.lift_comp] · intro q' hq' apply hc.algHom_ext' apply_fun LinearMap.restrictScalars R at hq' @@ -390,7 +390,7 @@ theorem Algebra.IsPushout.symm (h : Algebra.IsPushout R S R' S') : Algebra.IsPus (toAlgHom R S S').toLinearMap = (e.toLinearMap.restrictScalars R).comp (TensorProduct.mk R R' S 1) := by ext - simp [h.1.equiv_tmul, Algebra.smul_def] + simp [e, h.1.equiv_tmul, Algebra.smul_def] constructor rw [this] exact (TensorProduct.isBaseChange R S R').comp (IsBaseChange.ofEquiv e) diff --git a/Mathlib/RingTheory/Localization/BaseChange.lean b/Mathlib/RingTheory/Localization/BaseChange.lean index bbddee1b93312..6099cbf7f523d 100644 --- a/Mathlib/RingTheory/Localization/BaseChange.lean +++ b/Mathlib/RingTheory/Localization/BaseChange.lean @@ -129,5 +129,5 @@ instance (R M : Type*) [CommRing R] [AddCommGroup M] [Module R M] suffices (if a = b then f m else 0) = e (1 ⊗ₜ[R] if a = b then m else 0) by simpa [e', Finsupp.single_apply, -EmbeddingLike.apply_eq_iff_eq, apply_ite e] split_ifs with h - swap; · simp - simp [e, IsBaseChange.equiv_tmul] + · simp [e, IsBaseChange.equiv_tmul] + · simp only [tmul_zero, LinearEquiv.trans_apply, LinearEquiv.restrictScalars_apply, map_zero] diff --git a/Mathlib/RingTheory/Localization/Ideal.lean b/Mathlib/RingTheory/Localization/Ideal.lean index fb2341191e944..5e156755a12ed 100644 --- a/Mathlib/RingTheory/Localization/Ideal.lean +++ b/Mathlib/RingTheory/Localization/Ideal.lean @@ -47,7 +47,7 @@ private def map_ideal (I : Ideal R) : Ideal S where let Z : { x // x ∈ I } := ⟨(a'.2 : R) * (b'.1 : R) + (b'.2 : R) * (a'.1 : R), I.add_mem (I.mul_mem_left _ b'.1.2) (I.mul_mem_left _ a'.1.2)⟩ use ⟨Z, a'.2 * b'.2⟩ - simp only [RingHom.map_add, Submodule.coe_mk, Submonoid.coe_mul, RingHom.map_mul] + simp only [Z, RingHom.map_add, Submodule.coe_mk, Submonoid.coe_mul, RingHom.map_mul] rw [add_mul, ← mul_assoc a, ha, mul_comm (algebraMap R S a'.2) (algebraMap R S b'.2), ← mul_assoc b, hb] ring @@ -56,7 +56,7 @@ private def map_ideal (I : Ideal R) : Ideal S where obtain ⟨c', hc⟩ := IsLocalization.surj M c let Z : { x // x ∈ I } := ⟨c'.1 * x'.1, I.mul_mem_left c'.1 x'.1.2⟩ use ⟨Z, c'.2 * x'.2⟩ - simp only [← hx, ← hc, smul_eq_mul, Submodule.coe_mk, Submonoid.coe_mul, RingHom.map_mul] + simp only [Z, ← hx, ← hc, smul_eq_mul, Submodule.coe_mk, Submonoid.coe_mul, RingHom.map_mul] ring theorem mem_map_algebraMap_iff {I : Ideal R} {z} : z ∈ Ideal.map (algebraMap R S) I ↔ diff --git a/Mathlib/RingTheory/PowerSeries/Order.lean b/Mathlib/RingTheory/PowerSeries/Order.lean index 381b894105712..97eaca3724e77 100644 --- a/Mathlib/RingTheory/PowerSeries/Order.lean +++ b/Mathlib/RingTheory/PowerSeries/Order.lean @@ -329,9 +329,9 @@ theorem divided_by_X_pow_orderMul {f g : R⟦X⟧} (hf : f ≠ 0) (hg : g ≠ 0) divided_by_X_pow_order (mul_ne_zero hf hg) := by set df := f.order.lift (order_finite_iff_ne_zero.mpr hf) set dg := g.order.lift (order_finite_iff_ne_zero.mpr hg) - set dfg := (f * g).order.lift (order_finite_iff_ne_zero.mpr (mul_ne_zero hf hg)) with hdfg + set dfg := (f * g).order.lift (order_finite_iff_ne_zero.mpr (mul_ne_zero hf hg)) have H_add_d : df + dg = dfg := by - simp_all [order_mul f g] + simp_all [df, dg, dfg, order_mul f g] have H := self_eq_X_pow_order_mul_divided_by_X_pow_order (mul_ne_zero hf hg) have : f * g = X ^ dfg * (divided_by_X_pow_order hf * divided_by_X_pow_order hg) := by calc diff --git a/Mathlib/RingTheory/RingHom/FinitePresentation.lean b/Mathlib/RingTheory/RingHom/FinitePresentation.lean index 195c2d6891823..f2503b3b0b8f9 100644 --- a/Mathlib/RingTheory/RingHom/FinitePresentation.lean +++ b/Mathlib/RingTheory/RingHom/FinitePresentation.lean @@ -151,7 +151,7 @@ theorem finitePresentation_ofLocalizationSpanTarget : have : ∃ (a : S) (hb : a ∈ s), (Ideal.Quotient.mk I) (g' ⟨a, hb⟩) = g.val := by obtain ⟨g, hg⟩ := g convert hg - simp [t] + simp [A, f', t] obtain ⟨r, hr, hrr⟩ := this simp only [f'] rw [← hrr, Ideal.Quotient.liftₐ_apply, Ideal.Quotient.lift_mk] diff --git a/Mathlib/SetTheory/Ordinal/Notation.lean b/Mathlib/SetTheory/Ordinal/Notation.lean index 52e8d862840e0..9d8409124777a 100644 --- a/Mathlib/SetTheory/Ordinal/Notation.lean +++ b/Mathlib/SetTheory/Ordinal/Notation.lean @@ -831,7 +831,7 @@ theorem repr_opow_aux₂ {a0 a'} [N0 : NF a0] [Na' : NF a'] (m : ℕ) (d : ω by_cases h : m = 0 · simp only [R, R', h, ONote.ofNat, Nat.cast_zero, zero_add, ONote.repr, mul_zero, ONote.opowAux, add_zero] - · simp only [R', ONote.repr_scale, ONote.repr, ONote.mulNat_eq_mul, ONote.opowAux, + · simp only [α', ω0, R, 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 _ omega0_pos) diff --git a/Mathlib/Topology/Algebra/Group/OpenMapping.lean b/Mathlib/Topology/Algebra/Group/OpenMapping.lean index e3e52fc76e603..f2e6dce2efef1 100644 --- a/Mathlib/Topology/Algebra/Group/OpenMapping.lean +++ b/Mathlib/Topology/Algebra/Group/OpenMapping.lean @@ -114,6 +114,6 @@ theorem MonoidHom.isOpenMap_of_sigmaCompact let A : MulAction G H := MulAction.compHom _ f have : ContinuousSMul G H := continuousSMul_compHom h'f have : IsPretransitive G H := isPretransitive_compHom hf - have : f = (fun (g : G) ↦ g • (1 : H)) := by simp [MulAction.compHom_smul_def] + have : f = (fun (g : G) ↦ g • (1 : H)) := by simp [A, MulAction.compHom_smul_def] rw [this] exact isOpenMap_smul_of_sigmaCompact _ diff --git a/Mathlib/Topology/Algebra/Module/LinearPMap.lean b/Mathlib/Topology/Algebra/Module/LinearPMap.lean index dc97ddf2c8859..1fccf24ee319b 100644 --- a/Mathlib/Topology/Algebra/Module/LinearPMap.lean +++ b/Mathlib/Topology/Algebra/Module/LinearPMap.lean @@ -155,7 +155,7 @@ theorem closureHasCore (f : E →ₗ.[R] F) : f.closure.HasCore f.domain := by intro hx exact f.le_closure.1 hx let z : f.closure.domain := ⟨y.1, f.le_closure.1 y.2⟩ - have hyz : (y : E) = z := by simp + have hyz : (y : E) = z := by simp [z] rw [f.le_closure.2 hyz] exact domRestrict_apply (hxy.trans hyz) diff --git a/Mathlib/Topology/Algebra/ProperAction/Basic.lean b/Mathlib/Topology/Algebra/ProperAction/Basic.lean index 5b68b2d2d0d6f..12dc275ef0e31 100644 --- a/Mathlib/Topology/Algebra/ProperAction/Basic.lean +++ b/Mathlib/Topology/Algebra/ProperAction/Basic.lean @@ -129,15 +129,15 @@ theorem t2Space_of_properSMul_of_t2Group [h_proper : ProperSMul G X] [T2Space G] have proper_f : IsProperMap f := by refine IsClosedEmbedding.isProperMap ⟨?_, ?_⟩ · let g := fun gx : G × X ↦ gx.2 - have : Function.LeftInverse g f := fun x ↦ by simp + have : Function.LeftInverse g f := fun x ↦ by simp [f, g] exact this.isEmbedding (by fun_prop) (by fun_prop) - · have : range f = ({1} ×ˢ univ) := by simp + · have : range f = ({1} ×ˢ univ) := by simp [f] rw [this] exact isClosed_singleton.prod isClosed_univ rw [t2_iff_isClosed_diagonal] let g := fun gx : G × X ↦ (gx.1 • gx.2, gx.2) have proper_g : IsProperMap g := (properSMul_iff G X).1 h_proper - have : g ∘ f = fun x ↦ (x, x) := by ext x <;> simp + have : g ∘ f = fun x ↦ (x, x) := by ext x <;> simp [f, g] have range_gf : range (g ∘ f) = diagonal X := by simp [this] rw [← range_gf] exact (proper_f.comp proper_g).isClosed_range diff --git a/Mathlib/Topology/Category/Profinite/Nobeling.lean b/Mathlib/Topology/Category/Profinite/Nobeling.lean index 686efbf1e5992..c244872abc715 100644 --- a/Mathlib/Topology/Category/Profinite/Nobeling.lean +++ b/Mathlib/Topology/Category/Profinite/Nobeling.lean @@ -1342,7 +1342,7 @@ theorem CC_exact {f : LocallyConstant C ℤ} (hf : Linear_CC' C hsC ho f = 0) : cases' hx with hx₀ hx₁ · have hx₀' : ProjRestrict C (ord I · < o) ⟨x, hx⟩ = x := by simpa only [ProjRestrict, Set.MapsTo.val_restrict_apply] using C0_projOrd C hsC ho hx₀ - simp only [πs_apply_apply, hx₀', hx₀, LocallyConstant.piecewise'_apply_left, + simp only [C₀C, πs_apply_apply, hx₀', hx₀, LocallyConstant.piecewise'_apply_left, LocallyConstant.coe_comap, ContinuousMap.coe_mk, Function.comp_apply] · have hx₁' : (ProjRestrict C (ord I · < o) ⟨x, hx⟩).val ∈ π (C1 C ho) (ord I · < o) := by simpa only [ProjRestrict, Set.MapsTo.val_restrict_apply] using ⟨x, hx₁, rfl⟩ diff --git a/Mathlib/Topology/Category/TopCat/EffectiveEpi.lean b/Mathlib/Topology/Category/TopCat/EffectiveEpi.lean index 7502f5bebf01b..a42db3835f71f 100644 --- a/Mathlib/Topology/Category/TopCat/EffectiveEpi.lean +++ b/Mathlib/Topology/Category/TopCat/EffectiveEpi.lean @@ -69,7 +69,8 @@ theorem effectiveEpi_iff_isQuotientMap {B X : TopCat.{u}} (π : X ⟶ B) : /- The key to proving that the coequalizer has the quotient topology is `TopCat.coequalizer_isOpen_iff` which characterises the open sets in a coequalizer. -/ · ext U - have : π ≫ i.hom = colimit.ι F WalkingParallelPair.one := by simp [i, ← Iso.eq_comp_inv] + have : π ≫ i.hom = colimit.ι F WalkingParallelPair.one := by + simp [F, i, ← Iso.eq_comp_inv] rw [isOpen_coinduced (f := (homeoOfIso i ∘ π)), coequalizer_isOpen_iff _ U, ← this] rfl diff --git a/Mathlib/Topology/Connected/PathConnected.lean b/Mathlib/Topology/Connected/PathConnected.lean index 6f82df2380b21..47f59cf4a2bd6 100644 --- a/Mathlib/Topology/Connected/PathConnected.lean +++ b/Mathlib/Topology/Connected/PathConnected.lean @@ -544,7 +544,6 @@ def truncate {X : Type*} [TopologicalSpace X] {a b : X} (γ : Path a b) (t₀ t γ.continuous_extend.comp ((continuous_subtype_val.max continuous_const).min continuous_const) source' := by simp only [min_def, max_def'] - norm_cast split_ifs with h₁ h₂ h₃ h₄ · simp [γ.extend_of_le_zero h₁] · congr @@ -554,7 +553,6 @@ def truncate {X : Type*} [TopologicalSpace X] {a b : X} (γ : Path a b) (t₀ t all_goals rfl target' := by simp only [min_def, max_def'] - norm_cast split_ifs with h₁ h₂ h₃ · simp [γ.extend_of_one_le h₂] · rfl diff --git a/Mathlib/Topology/ContinuousMap/Algebra.lean b/Mathlib/Topology/ContinuousMap/Algebra.lean index 42a37d593e9a0..095d12e987b09 100644 --- a/Mathlib/Topology/ContinuousMap/Algebra.lean +++ b/Mathlib/Topology/ContinuousMap/Algebra.lean @@ -760,8 +760,8 @@ theorem Subalgebra.SeparatesPoints.strongly {s : Subalgebra 𝕜 C(α, 𝕜)} (h let f' : s := ((b - a) * (f x - f y)⁻¹) • (algebraMap _ s (f x) - (⟨f, hf⟩ : s)) + algebraMap _ s a refine ⟨f', f'.prop, ?_, ?_⟩ - · simp [f'] - · simp [f', inv_mul_cancel_right₀ hxy] + · simp [a, b, f'] + · simp [a, b, f', inv_mul_cancel_right₀ hxy] end ContinuousMap diff --git a/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean b/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean index b3ec116499bc4..e62e6489c8c9e 100644 --- a/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean +++ b/Mathlib/Topology/ContinuousMap/StoneWeierstrass.lean @@ -213,10 +213,10 @@ theorem sublattice_closure_eq_top (L : Set C(X, ℝ)) (nA : L.Nonempty) have lt_h : ∀ x z, f z - ε < (h x : X → ℝ) z := by intro x z obtain ⟨y, ym, zm⟩ := Set.exists_set_mem_of_union_eq_top _ _ (ys_w x) z - dsimp + dsimp [h] simp only [Subtype.coe_mk, coe_sup', Finset.sup'_apply, Finset.lt_sup'_iff] exact ⟨y, ym, zm⟩ - have h_eq : ∀ x, (h x : X → ℝ) x = f x := by intro x; simp [w₁] + have h_eq : ∀ x, (h x : X → ℝ) x = f x := by intro x; simp [h, w₁] -- For each `x`, we define `W x` to be `{z | h x z < f z + ε}`, let W : X → Set X := fun x => {z | (h x : X → ℝ) z < f z + ε} -- This is still a neighbourhood of `x`. @@ -247,10 +247,10 @@ theorem sublattice_closure_eq_top (L : Set C(X, ℝ)) (nA : L.Nonempty) intros; simp only [← Metric.mem_ball, Real.ball_eq_Ioo, Set.mem_Ioo, and_comm]] fconstructor · dsimp - simp only [Finset.inf'_lt_iff, ContinuousMap.inf'_apply] + simp only [k, Finset.inf'_lt_iff, ContinuousMap.inf'_apply] exact Set.exists_set_mem_of_union_eq_top _ _ xs_w z · dsimp - simp only [Finset.lt_inf'_iff, ContinuousMap.inf'_apply] + simp only [k, Finset.lt_inf'_iff, ContinuousMap.inf'_apply] rintro x - apply lt_h diff --git a/Mathlib/Topology/FiberBundle/Constructions.lean b/Mathlib/Topology/FiberBundle/Constructions.lean index 21d63b4d3c851..a123634cbb964 100644 --- a/Mathlib/Topology/FiberBundle/Constructions.lean +++ b/Mathlib/Topology/FiberBundle/Constructions.lean @@ -148,7 +148,7 @@ theorem Prod.continuous_to_fun : ContinuousOn (Prod.toFun' e₁ e₂) · rw [e₁.source_eq, e₂.source_eq] exact mapsTo_preimage _ _ rintro ⟨b, v₁, v₂⟩ ⟨hb₁, _⟩ - simp only [f₃, Prod.toFun', Prod.mk.inj_iff, Function.comp_apply, and_true] + simp only [f₁, f₂, f₃, Prod.toFun', Prod.mk.inj_iff, Function.comp_apply, and_true] rw [e₁.coe_fst] rw [e₁.source_eq, mem_preimage] exact hb₁ diff --git a/Mathlib/Topology/Sheaves/SheafCondition/PairwiseIntersections.lean b/Mathlib/Topology/Sheaves/SheafCondition/PairwiseIntersections.lean index 196cd67b22551..0ad5004a30a59 100644 --- a/Mathlib/Topology/Sheaves/SheafCondition/PairwiseIntersections.lean +++ b/Mathlib/Topology/Sheaves/SheafCondition/PairwiseIntersections.lean @@ -344,7 +344,7 @@ def interUnionPullbackConeLift : s.pt ⟶ F.1.obj (op (U ⊔ V)) := by rcases j with (⟨⟨_ | _⟩⟩ | ⟨⟨_ | _⟩, ⟨_⟩⟩) <;> rcases g with ⟨⟩ <;> dsimp [Pairwise.diagram] <;> - simp only [Category.id_comp, s.condition, CategoryTheory.Functor.map_id, Category.comp_id] + simp only [ι, Category.id_comp, s.condition, CategoryTheory.Functor.map_id, Category.comp_id] rw [← cancel_mono (F.1.map (eqToHom <| inf_comm U V : U ⊓ V ⟶ _).op), Category.assoc, Category.assoc, ← F.1.map_comp, ← F.1.map_comp] exact s.condition.symm diff --git a/Mathlib/Topology/UniformSpace/Completion.lean b/Mathlib/Topology/UniformSpace/Completion.lean index bfea3bac43b00..7e1f98e7b1de7 100644 --- a/Mathlib/Topology/UniformSpace/Completion.lean +++ b/Mathlib/Topology/UniformSpace/Completion.lean @@ -80,7 +80,7 @@ private theorem symm_gen : map Prod.swap ((𝓤 α).lift' gen) ≤ (𝓤 α).lif { p : CauchyFilter α × CauchyFilter α | s ∈ (p.2.val ×ˢ p.1.val : Filter (α × α)) } have h₁ : map Prod.swap ((𝓤 α).lift' gen) = (𝓤 α).lift' f := by delta gen - simp [map_lift'_eq, monotone_setOf, Filter.monotone_mem, Function.comp_def, + simp [f, map_lift'_eq, monotone_setOf, Filter.monotone_mem, Function.comp_def, image_swap_eq_preimage_swap] have h₂ : (𝓤 α).lift' f ≤ (𝓤 α).lift' gen := uniformity_lift_le_swap From 128f5afc197d4518cb5fd4f1a073e60160aa3c5b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Javier=20L=C3=B3pez-Contreras?= <“jlopezcontreras10@gmail.com”> Date: Tue, 10 Dec 2024 13:44:15 +0000 Subject: [PATCH 633/829] refactor(LocalRing/Basic): generalize to non-commutative (semi)rings (#18652) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Small generalization of some Lemmas in LocalRing.Basic Co-authored-by: Yaël Dillies --- Mathlib/RingTheory/LocalRing/Basic.lean | 33 ++++++++++++++++--------- 1 file changed, 22 insertions(+), 11 deletions(-) diff --git a/Mathlib/RingTheory/LocalRing/Basic.lean b/Mathlib/RingTheory/LocalRing/Basic.lean index ff066f8280c34..48a98a6d9adb0 100644 --- a/Mathlib/RingTheory/LocalRing/Basic.lean +++ b/Mathlib/RingTheory/LocalRing/Basic.lean @@ -16,11 +16,11 @@ We prove basic properties of local rings. variable {R S : Type*} -section CommSemiring +namespace IsLocalRing -variable [CommSemiring R] +section Semiring -namespace IsLocalRing +variable [Semiring R] theorem of_isUnit_or_isUnit_of_isUnit_add [Nontrivial R] (h : ∀ a b : R, IsUnit (a + b) → IsUnit a ∨ IsUnit b) : IsLocalRing R := @@ -28,8 +28,15 @@ theorem of_isUnit_or_isUnit_of_isUnit_add [Nontrivial R] /-- A semiring is local if it is nontrivial and the set of nonunits is closed under the addition. -/ theorem of_nonunits_add [Nontrivial R] - (h : ∀ a b : R, a ∈ nonunits R → b ∈ nonunits R → a + b ∈ nonunits R) : IsLocalRing R := - ⟨fun {a b} hab => or_iff_not_and_not.2 fun H => h a b H.1 H.2 <| hab.symm ▸ isUnit_one⟩ + (h : ∀ a b : R, a ∈ nonunits R → b ∈ nonunits R → a + b ∈ nonunits R) : IsLocalRing R where + isUnit_or_isUnit_of_add_one {a b} hab := + or_iff_not_and_not.2 fun H => h a b H.1 H.2 <| hab.symm ▸ isUnit_one + +end Semiring + +section CommSemiring + +variable [CommSemiring R] /-- A semiring is local if it has a unique maximal ideal. -/ theorem of_unique_max_ideal (h : ∃! I : Ideal R, I.IsMaximal) : IsLocalRing R := @@ -65,8 +72,6 @@ theorem isUnit_or_isUnit_of_isUnit_add {a b : R} (h : IsUnit (a + b)) : IsUnit a theorem nonunits_add {a b : R} (ha : a ∈ nonunits R) (hb : b ∈ nonunits R) : a + b ∈ nonunits R := fun H => not_or_intro ha hb (isUnit_or_isUnit_of_isUnit_add H) -end IsLocalRing - @[deprecated (since := "2024-11-11")] alias LocalRing.of_isUnit_or_isUnit_of_isUnit_add := IsLocalRing.of_isUnit_or_isUnit_of_isUnit_add @@ -87,15 +92,19 @@ alias LocalRing.nonunits_add := IsLocalRing.nonunits_add end CommSemiring -namespace IsLocalRing +section Ring -variable [CommRing R] +variable [Ring R] theorem of_isUnit_or_isUnit_one_sub_self [Nontrivial R] (h : ∀ a : R, IsUnit a ∨ IsUnit (1 - a)) : IsLocalRing R := ⟨fun {a b} hab => add_sub_cancel_left a b ▸ hab.symm ▸ h a⟩ -variable [IsLocalRing R] +end Ring + +section CommRing + +variable [CommRing R] [IsLocalRing R] theorem isUnit_or_isUnit_one_sub_self (a : R) : IsUnit a ∨ IsUnit (1 - a) := isUnit_or_isUnit_of_isUnit_add <| (add_sub_cancel a 1).symm ▸ isUnit_one @@ -106,7 +115,7 @@ theorem isUnit_of_mem_nonunits_one_sub_self (a : R) (h : 1 - a ∈ nonunits R) : theorem isUnit_one_sub_self_of_mem_nonunits (a : R) (h : a ∈ nonunits R) : IsUnit (1 - a) := or_iff_not_imp_left.1 (isUnit_or_isUnit_one_sub_self a) h -theorem of_surjective' [CommRing S] [Nontrivial S] (f : R →+* S) (hf : Function.Surjective f) : +theorem of_surjective' [Ring S] [Nontrivial S] (f : R →+* S) (hf : Function.Surjective f) : IsLocalRing S := of_isUnit_or_isUnit_one_sub_self (by intro b @@ -132,6 +141,8 @@ alias LocalRing.isUnit_one_sub_self_of_mem_nonunits := @[deprecated (since := "2024-11-11")] alias LocalRing.of_surjective' := IsLocalRing.of_surjective' +end CommRing + end IsLocalRing namespace Field From 965bd1ea790c952273326d1c99f55f88114de87b Mon Sep 17 00:00:00 2001 From: sgouezel Date: Tue, 10 Dec 2024 14:12:18 +0000 Subject: [PATCH 634/829] feat: the iterated derivative of an analytic function (#19862) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Prove the formula `D^n f (x ; v_1, ..., v_n) = ∑ pₙ (v_{σ (1)}, ..., v_{σ (n)})` if `pₙ` is the `n`-th term in a power series expansion of `f`, where the sum is over all permutations of `Fin n`. Deduce that the iterated derivative is symmetric. --- Mathlib.lean | 1 + Mathlib/Analysis/Analytic/IteratedFDeriv.lean | 277 ++++++++++++++++++ .../Analysis/Calculus/FDeriv/Analytic.lean | 4 + 3 files changed, 282 insertions(+) create mode 100644 Mathlib/Analysis/Analytic/IteratedFDeriv.lean diff --git a/Mathlib.lean b/Mathlib.lean index 9d04fd689ce22..b5020b3a7fbb0 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1072,6 +1072,7 @@ import Mathlib.Analysis.Analytic.Composition import Mathlib.Analysis.Analytic.Constructions import Mathlib.Analysis.Analytic.Inverse import Mathlib.Analysis.Analytic.IsolatedZeros +import Mathlib.Analysis.Analytic.IteratedFDeriv import Mathlib.Analysis.Analytic.Linear import Mathlib.Analysis.Analytic.Meromorphic import Mathlib.Analysis.Analytic.OfScalars diff --git a/Mathlib/Analysis/Analytic/IteratedFDeriv.lean b/Mathlib/Analysis/Analytic/IteratedFDeriv.lean new file mode 100644 index 0000000000000..b4c95ff5110a4 --- /dev/null +++ b/Mathlib/Analysis/Analytic/IteratedFDeriv.lean @@ -0,0 +1,277 @@ +/- +Copyright (c) 2024 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel +-/ +import Mathlib.Analysis.Calculus.ContDiff.Basic +import Mathlib.Analysis.Calculus.ContDiff.CPolynomial +import Mathlib.Data.Fintype.Perm + +/-! +# The iterated derivative of an analytic function + +If a function is analytic, written as `f (x + y) = ∑ pₙ (y, ..., y)` then its `n`-th iterated +derivative at `x` is given by `(v₁, ..., vₙ) ↦ ∑ pₙ (v_{σ (1)}, ..., v_{σ (n)})` where the sum +is over all permutations of `{1, ..., n}`. In particular, it is symmetric. + +This generalizes the result of `HasFPowerSeriesOnBall.factorial_smul` giving +`D^n f (v, ..., v) = n! * pₙ (v, ..., v)`. + +## Main result + +* `HasFPowerSeriesOnBall.iteratedFDeriv_eq_sum` shows that + `iteratedFDeriv 𝕜 n f x v = ∑ σ : Perm (Fin n), p n (fun i ↦ v (σ i))`, + when `f` has `p` as power series within the set `s` on the ball `B (x, r)`. +* `ContDiffAt.iteratedFDeriv_comp_perm` proves the symmetry of the iterated derivative of an + analytic function, in the form `iteratedFDeriv 𝕜 n f x (v ∘ σ) = iteratedFDeriv 𝕜 n f x v` + for any permutation `σ` of `Fin n`. + +Versions within sets are also given. + +## Implementation + +To prove the formula for the iterated derivative, we decompose an analytic function as +the sum of `fun y ↦ pₙ (y, ..., y)` and the rest. For the former, its iterated derivative follows +from the formula for iterated derivatives of multilinear maps +(see `ContinuousMultilinearMap.iteratedFDeriv_comp_diagonal`). For the latter, we show by +induction on `n` that if the `n`-th term in a power series is zero, then the `n`-th iterated +derivative vanishes (see `HasFPowerSeriesWithinOnBall.iteratedFDerivWithin_eq_zero`). + +All these results are proved assuming additionally that the function is analytic on the relevant +set (which does not follow from the fact that the function has a power series, if the target space +is not complete). This makes it possible to avoid all completeness assumptions in the final +statements. When needed, we give versions of some statements assuming completeness and dropping +analyticity, for ease of use. +-/ + +open scoped ENNReal Topology ContDiff +open Equiv Set + +variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] + {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] + {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] + {f : E → F} {p : FormalMultilinearSeries 𝕜 E F} {s : Set E} {x : E} {r : ℝ≥0∞} + +/-- Formal multilinear series associated to the iterated derivative, defined by iterating +`p ↦ p.derivSeries` and currying suitably. It is defined so that, if a function has `p` as a power +series, then its iterated derivative of order `k` has `p.iteratedFDerivSeries k` as a power +series. -/ +noncomputable def FormalMultilinearSeries.iteratedFDerivSeries + (p : FormalMultilinearSeries 𝕜 E F) (k : ℕ) : + FormalMultilinearSeries 𝕜 E (E [×k]→L[𝕜] F) := + match k with + | 0 => (continuousMultilinearCurryFin0 𝕜 E F).symm + |>.toContinuousLinearEquiv.toContinuousLinearMap.compFormalMultilinearSeries p + | (k + 1) => (continuousMultilinearCurryLeftEquiv 𝕜 (fun _ : Fin (k + 1) ↦ E) F).symm + |>.toContinuousLinearEquiv.toContinuousLinearMap.compFormalMultilinearSeries + (p.iteratedFDerivSeries k).derivSeries + +/-- If a function has a power series on a ball, then so do its iterated derivatives. -/ +protected theorem HasFPowerSeriesWithinOnBall.iteratedFDerivWithin + (h : HasFPowerSeriesWithinOnBall f p s x r) (h' : AnalyticOn 𝕜 f s) + (k : ℕ) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) : + HasFPowerSeriesWithinOnBall (iteratedFDerivWithin 𝕜 k f s) + (p.iteratedFDerivSeries k) s x r := by + induction k with + | zero => + exact (continuousMultilinearCurryFin0 𝕜 E F).symm + |>.toContinuousLinearEquiv.toContinuousLinearMap.comp_hasFPowerSeriesWithinOnBall h + | succ k ih => + rw [iteratedFDerivWithin_succ_eq_comp_left] + apply (continuousMultilinearCurryLeftEquiv 𝕜 (fun _ : Fin (k + 1) ↦ E) F).symm + |>.toContinuousLinearEquiv.toContinuousLinearMap.comp_hasFPowerSeriesWithinOnBall + (ih.fderivWithin_of_mem_of_analyticOn (h'.iteratedFDerivWithin hs _) hs hx) + +lemma FormalMultilinearSeries.iteratedFDerivSeries_eq_zero {k n : ℕ} + (h : p (n + k) = 0) : p.iteratedFDerivSeries k n = 0 := by + induction k generalizing n with + | zero => + ext + have : p n = 0 := p.congr_zero rfl h + simp [FormalMultilinearSeries.iteratedFDerivSeries, this] + | succ k ih => + ext + simp only [iteratedFDerivSeries, Nat.succ_eq_add_one, + ContinuousLinearMap.compFormalMultilinearSeries_apply, + ContinuousLinearMap.compContinuousMultilinearMap_coe, ContinuousLinearEquiv.coe_coe, + LinearIsometryEquiv.coe_toContinuousLinearEquiv, Function.comp_apply, + continuousMultilinearCurryLeftEquiv_symm_apply, ContinuousMultilinearMap.zero_apply] + rw [derivSeries_eq_zero] + · rfl + · apply ih + apply p.congr_zero (by abel) h + +/-- If the `n`-th term in a power series is zero, then the `n`-th derivative of the corresponding +function vanishes. -/ +lemma HasFPowerSeriesWithinOnBall.iteratedFDerivWithin_eq_zero + (h : HasFPowerSeriesWithinOnBall f p s x r) (h' : AnalyticOn 𝕜 f s) + (hu : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {n : ℕ} (hn : p n = 0) : + iteratedFDerivWithin 𝕜 n f s x = 0 := by + have : iteratedFDerivWithin 𝕜 n f s x = p.iteratedFDerivSeries n 0 (fun _ ↦ 0) := + ((h.iteratedFDerivWithin h' n hu hx).coeff_zero _).symm + rw [this, p.iteratedFDerivSeries_eq_zero (p.congr_zero (Nat.zero_add n).symm hn)] + rfl + +lemma ContinuousMultilinearMap.iteratedFDeriv_comp_diagonal + {n : ℕ} (f : E [×n]→L[𝕜] F) (x : E) (v : Fin n → E) : + iteratedFDeriv 𝕜 n (fun x ↦ f (fun _ ↦ x)) x v = ∑ σ : Perm (Fin n), f (fun i ↦ v (σ i)) := by + rw [← sum_comp (Equiv.inv (Perm (Fin n)))] + let g : E →L[𝕜] (Fin n → E) := ContinuousLinearMap.pi (fun i ↦ ContinuousLinearMap.id 𝕜 E) + change iteratedFDeriv 𝕜 n (f ∘ g) x v = _ + rw [ContinuousLinearMap.iteratedFDeriv_comp_right _ f.contDiff _ le_rfl, f.iteratedFDeriv_eq] + simp only [ContinuousMultilinearMap.iteratedFDeriv, + ContinuousMultilinearMap.compContinuousLinearMap_apply, ContinuousMultilinearMap.sum_apply, + ContinuousMultilinearMap.iteratedFDerivComponent_apply, Set.mem_range, Pi.compRightL_apply] + rw [← sum_comp (Equiv.embeddingEquivOfFinite (Fin n))] + congr with σ + congr with i + have A : ∃ y, σ y = i := by + have : Function.Bijective σ := (Fintype.bijective_iff_injective_and_card _).2 ⟨σ.injective, rfl⟩ + exact this.surjective i + rcases A with ⟨y, rfl⟩ + simp only [EmbeddingLike.apply_eq_iff_eq, exists_eq, ↓reduceDIte, + Function.Embedding.toEquivRange_symm_apply_self, ContinuousLinearMap.coe_pi', + ContinuousLinearMap.coe_id', id_eq, g] + congr 1 + symm + simp [coe_fn_mk, inv_apply, Perm.inv_def, + ofBijective_symm_apply_apply, Function.Embedding.equivOfFiniteSelfEmbedding] + +private lemma HasFPowerSeriesWithinOnBall.iteratedFDerivWithin_eq_sum_of_subset + (h : HasFPowerSeriesWithinOnBall f p s x r) (h' : AnalyticOn 𝕜 f s) + (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) + {n : ℕ} (v : Fin n → E) (h's : s ⊆ EMetric.ball x r) : + iteratedFDerivWithin 𝕜 n f s x v = ∑ σ : Perm (Fin n), p n (fun i ↦ v (σ i)) := by + have I : insert x s ∩ EMetric.ball x r = s := by + rw [Set.insert_eq_of_mem hx] + exact Set.inter_eq_left.2 h's + have fcont : ContDiffOn 𝕜 (↑n) f s := by + apply AnalyticOn.contDiffOn _ hs + simpa [I] using h' + let g : E → F := fun z ↦ p n (fun _ ↦ z - x) + have gcont : ContDiff 𝕜 ω g := by + apply (p n).contDiff.comp + exact contDiff_pi.2 (fun i ↦ contDiff_id.sub contDiff_const) + let q : FormalMultilinearSeries 𝕜 E F := fun k ↦ if h : n = k then (h ▸ p n) else 0 + have A : HasFiniteFPowerSeriesOnBall g q x (n + 1) r := by + apply HasFiniteFPowerSeriesOnBall.mk' _ h.r_pos + · intro y hy + rw [Finset.sum_eq_single_of_mem n] + · simp [q, g] + · simp + · intro i hi h'i + simp [q, h'i.symm] + · intro m hm + have : n ≠ m := by omega + simp [q, this] + have B : HasFPowerSeriesWithinOnBall g q s x r := + A.toHasFPowerSeriesOnBall.hasFPowerSeriesWithinOnBall + have J1 : iteratedFDerivWithin 𝕜 n f s x = + iteratedFDerivWithin 𝕜 n g s x + iteratedFDerivWithin 𝕜 n (f - g) s x := by + have : f = g + (f - g) := by abel + nth_rewrite 1 [this] + rw [iteratedFDerivWithin_add_apply (gcont.of_le le_top).contDiffOn + (by exact fcont.sub (gcont.of_le le_top).contDiffOn) hs hx] + have J2 : iteratedFDerivWithin 𝕜 n (f - g) s x = 0 := by + apply (h.sub B).iteratedFDerivWithin_eq_zero (h'.sub ?_) hs hx + · simp [q] + · apply gcont.contDiffOn.analyticOn + have J3 : iteratedFDerivWithin 𝕜 n g s x = iteratedFDeriv 𝕜 n g x := + iteratedFDerivWithin_eq_iteratedFDeriv hs (gcont.of_le le_top).contDiffAt hx + simp only [J1, J3, J2, add_zero] + let g' : E → F := fun z ↦ p n (fun _ ↦ z) + have : g = fun z ↦ g' (z - x) := rfl + rw [this, iteratedFDeriv_comp_sub] + exact (p n).iteratedFDeriv_comp_diagonal _ v + +/-- If a function has a power series in a ball, then its `n`-th iterated derivative is given by +`(v₁, ..., vₙ) ↦ ∑ pₙ (v_{σ (1)}, ..., v_{σ (n)})` where the sum is over all +permutations of `{1, ..., n}`.-/ +theorem HasFPowerSeriesWithinOnBall.iteratedFDerivWithin_eq_sum + (h : HasFPowerSeriesWithinOnBall f p s x r) (h' : AnalyticOn 𝕜 f s) + (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {n : ℕ} (v : Fin n → E) : + iteratedFDerivWithin 𝕜 n f s x v = ∑ σ : Perm (Fin n), p n (fun i ↦ v (σ i)) := by + have : iteratedFDerivWithin 𝕜 n f s x + = iteratedFDerivWithin 𝕜 n f (s ∩ EMetric.ball x r) x := + (iteratedFDerivWithin_inter_open EMetric.isOpen_ball (EMetric.mem_ball_self h.r_pos)).symm + rw [this] + apply HasFPowerSeriesWithinOnBall.iteratedFDerivWithin_eq_sum_of_subset + · exact h.mono inter_subset_left + · exact h'.mono inter_subset_left + · exact hs.inter EMetric.isOpen_ball + · exact ⟨hx, EMetric.mem_ball_self h.r_pos⟩ + · exact inter_subset_right + +/-- If a function has a power series in a ball, then its `n`-th iterated derivative is given by +`(v₁, ..., vₙ) ↦ ∑ pₙ (v_{σ (1)}, ..., v_{σ (n)})` where the sum is over all +permutations of `{1, ..., n}`.-/ +theorem HasFPowerSeriesOnBall.iteratedFDeriv_eq_sum + (h : HasFPowerSeriesOnBall f p x r) (h' : AnalyticOn 𝕜 f univ) {n : ℕ} (v : Fin n → E) : + iteratedFDeriv 𝕜 n f x v = ∑ σ : Perm (Fin n), p n (fun i ↦ v (σ i)) := by + simp only [← iteratedFDerivWithin_univ, ← hasFPowerSeriesWithinOnBall_univ] at h ⊢ + exact h.iteratedFDerivWithin_eq_sum h' uniqueDiffOn_univ (mem_univ x) v + +/-- If a function has a power series in a ball, then its `n`-th iterated derivative is given by +`(v₁, ..., vₙ) ↦ ∑ pₙ (v_{σ (1)}, ..., v_{σ (n)})` where the sum is over all +permutations of `{1, ..., n}`.-/ +theorem HasFPowerSeriesWithinOnBall.iteratedFDerivWithin_eq_sum_of_completeSpace [CompleteSpace F] + (h : HasFPowerSeriesWithinOnBall f p s x r) + (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {n : ℕ} (v : Fin n → E) : + iteratedFDerivWithin 𝕜 n f s x v = ∑ σ : Perm (Fin n), p n (fun i ↦ v (σ i)) := by + have : iteratedFDerivWithin 𝕜 n f s x + = iteratedFDerivWithin 𝕜 n f (s ∩ EMetric.ball x r) x := + (iteratedFDerivWithin_inter_open EMetric.isOpen_ball (EMetric.mem_ball_self h.r_pos)).symm + rw [this] + apply HasFPowerSeriesWithinOnBall.iteratedFDerivWithin_eq_sum_of_subset + · exact h.mono inter_subset_left + · apply h.analyticOn.mono + rw [insert_eq_of_mem hx] + · exact hs.inter EMetric.isOpen_ball + · exact ⟨hx, EMetric.mem_ball_self h.r_pos⟩ + · exact inter_subset_right + +/-- If a function has a power series in a ball, then its `n`-th iterated derivative is given by +`(v₁, ..., vₙ) ↦ ∑ pₙ (v_{σ (1)}, ..., v_{σ (n)})` where the sum is over all +permutations of `{1, ..., n}`.-/ +theorem HasFPowerSeriesOnBall.iteratedFDeriv_eq_sum_of_completeSpace [CompleteSpace F] + (h : HasFPowerSeriesOnBall f p x r) {n : ℕ} (v : Fin n → E) : + iteratedFDeriv 𝕜 n f x v = ∑ σ : Perm (Fin n), p n (fun i ↦ v (σ i)) := by + simp only [← iteratedFDerivWithin_univ, ← hasFPowerSeriesWithinOnBall_univ] at h ⊢ + exact h.iteratedFDerivWithin_eq_sum_of_completeSpace uniqueDiffOn_univ (mem_univ _) v + +/-- The `n`-th iterated derivative of an analytic function on a set is symmetric. -/ +theorem AnalyticOn.iteratedFDerivWithin_comp_perm + (h : AnalyticOn 𝕜 f s) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {n : ℕ} (v : Fin n → E) + (σ : Perm (Fin n)) : + iteratedFDerivWithin 𝕜 n f s x (v ∘ σ) = iteratedFDerivWithin 𝕜 n f s x v := by + rcases h x hx with ⟨p, r, hp⟩ + rw [hp.iteratedFDerivWithin_eq_sum h hs hx, hp.iteratedFDerivWithin_eq_sum h hs hx] + let e := Equiv.mulLeft σ + conv_rhs => rw [← Equiv.sum_comp e] + rfl + +/-- The `n`-th iterated derivative of an analytic function on a set is symmetric. -/ +theorem ContDiffWithinAt.iteratedFDerivWithin_comp_perm + (h : ContDiffWithinAt 𝕜 ω f s x) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {n : ℕ} (v : Fin n → E) + (σ : Perm (Fin n)) : + iteratedFDerivWithin 𝕜 n f s x (v ∘ σ) = iteratedFDerivWithin 𝕜 n f s x v := by + rcases h.contDiffOn' le_rfl (by simp) with ⟨u, u_open, xu, hu⟩ + rw [insert_eq_of_mem hx] at hu + have : iteratedFDerivWithin 𝕜 n f (s ∩ u) x = iteratedFDerivWithin 𝕜 n f s x := + iteratedFDerivWithin_inter_open u_open xu + rw [← this] + exact AnalyticOn.iteratedFDerivWithin_comp_perm hu.analyticOn (hs.inter u_open) ⟨hx, xu⟩ _ _ + +/-- The `n`-th iterated derivative of an analytic function is symmetric. -/ +theorem AnalyticOn.iteratedFDeriv_comp_perm + (h : AnalyticOn 𝕜 f univ) {n : ℕ} (v : Fin n → E) (σ : Perm (Fin n)) : + iteratedFDeriv 𝕜 n f x (v ∘ σ) = iteratedFDeriv 𝕜 n f x v := by + rw [← iteratedFDerivWithin_univ] + exact h.iteratedFDerivWithin_comp_perm uniqueDiffOn_univ (mem_univ x) _ _ + +/-- The `n`-th iterated derivative of an analytic function is symmetric. -/ +theorem ContDiffAt.iteratedFDeriv_comp_perm + (h : ContDiffAt 𝕜 ω f x) {n : ℕ} (v : Fin n → E) (σ : Perm (Fin n)) : + iteratedFDeriv 𝕜 n f x (v ∘ σ) = iteratedFDeriv 𝕜 n f x v := by + rw [← iteratedFDerivWithin_univ] + exact h.iteratedFDerivWithin_comp_perm uniqueDiffOn_univ (mem_univ x) _ _ diff --git a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean index d7b62c6e50328..fb89da18adf7c 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean @@ -762,6 +762,10 @@ private theorem factorial_smul' {n : ℕ} : ∀ {F : Type max u v} [NormedAddCom variable [CompleteSpace F] include h +/-- The iterated derivative of an analytic function, on vectors `(y, ..., y)`, is given by `n!` +times the `n`-th term in the power series. For a more general result giving the full iterated +derivative as a sum over the permutations of `Fin n`, see +`HasFPowerSeriesOnBall.iteratedFDeriv_eq_sum`. -/ theorem factorial_smul (n : ℕ) : n ! • p n (fun _ ↦ y) = iteratedFDeriv 𝕜 n f x (fun _ ↦ y) := by cases n From b7e2717540e62e79d1002d74437c791524233593 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 10 Dec 2024 14:59:53 +0000 Subject: [PATCH 635/829] chore: rename `forall_image2_iff` to `forall_mem_image2` (#19401) This matches the existing `image` lemmas. Also add the missing `exists` lemma. From GrowthInGroups (LeanCamCombi) --- .../Algebra/Order/Group/Pointwise/Bounds.lean | 2 +- Mathlib/Combinatorics/SetFamily/Shatter.lean | 2 +- Mathlib/Data/Finset/Image.lean | 6 +++-- Mathlib/Data/Finset/NAry.lean | 26 ++++++++++++------- Mathlib/Data/Finset/Sups.lean | 6 ++--- Mathlib/Data/Set/NAry.lean | 14 ++++++---- Mathlib/Data/Set/Sups.lean | 4 +-- Mathlib/Dynamics/OmegaLimit.lean | 2 +- Mathlib/NumberTheory/MaricaSchoenheim.lean | 4 +-- Mathlib/Order/Bounds/Image.lean | 16 ++++++------ .../ConditionallyCompleteLattice/Basic.lean | 2 +- Mathlib/Topology/CompactOpen.lean | 4 +-- Mathlib/Topology/EMetricSpace/Lipschitz.lean | 2 +- 13 files changed, 51 insertions(+), 39 deletions(-) diff --git a/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean b/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean index 294f6be7c92df..90e52eaed78c7 100644 --- a/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean +++ b/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean @@ -26,7 +26,7 @@ variable [Mul M] [Preorder M] [MulLeftMono M] @[to_additive] lemma mul_mem_upperBounds_mul (ha : a ∈ upperBounds s) (hb : b ∈ upperBounds t) : - a * b ∈ upperBounds (s * t) := forall_image2_iff.2 fun _ hx _ hy => mul_le_mul' (ha hx) (hb hy) + a * b ∈ upperBounds (s * t) := forall_mem_image2.2 fun _ hx _ hy => mul_le_mul' (ha hx) (hb hy) @[to_additive] lemma subset_upperBounds_mul (s t : Set M) : upperBounds s * upperBounds t ⊆ upperBounds (s * t) := diff --git a/Mathlib/Combinatorics/SetFamily/Shatter.lean b/Mathlib/Combinatorics/SetFamily/Shatter.lean index 27a76906f15d6..a9fb7d184ba69 100644 --- a/Mathlib/Combinatorics/SetFamily/Shatter.lean +++ b/Mathlib/Combinatorics/SetFamily/Shatter.lean @@ -137,7 +137,7 @@ lemma card_le_card_shatterer (𝒜 : Finset (Finset α)) : #𝒜 ≤ #𝒜.shatt rw [mem_memberSubfamily] at hv rw [← singleton_subset_iff (a := a), ← hsv] at hv exact hv.2 inter_subset_right - · refine forall_image.2 fun s hs ↦ mem_shatterer.2 fun t ht ↦ ?_ + · refine forall_mem_image.2 fun s hs ↦ mem_shatterer.2 fun t ht ↦ ?_ simp only [mem_inter, mem_shatterer] at hs rw [subset_insert_iff] at ht by_cases ha : a ∈ t diff --git a/Mathlib/Data/Finset/Image.lean b/Mathlib/Data/Finset/Image.lean index 5b0af20456dff..f5732b4cbc6e6 100644 --- a/Mathlib/Data/Finset/Image.lean +++ b/Mathlib/Data/Finset/Image.lean @@ -313,8 +313,10 @@ theorem mem_image : b ∈ s.image f ↔ ∃ a ∈ s, f a = b := by theorem mem_image_of_mem (f : α → β) {a} (h : a ∈ s) : f a ∈ s.image f := mem_image.2 ⟨_, h, rfl⟩ -theorem forall_image {p : β → Prop} : (∀ b ∈ s.image f, p b) ↔ ∀ a ∈ s, p (f a) := by - simp only [mem_image, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂] +lemma forall_mem_image {p : β → Prop} : (∀ y ∈ s.image f, p y) ↔ ∀ ⦃x⦄, x ∈ s → p (f x) := by simp +lemma exists_mem_image {p : β → Prop} : (∃ y ∈ s.image f, p y) ↔ ∃ x ∈ s, p (f x) := by simp + +@[deprecated (since := "2024-11-23")] alias forall_image := forall_mem_image theorem map_eq_image (f : α ↪ β) (s : Finset α) : s.map f = s.image f := eq_of_veq (s.map f).2.dedup.symm diff --git a/Mathlib/Data/Finset/NAry.lean b/Mathlib/Data/Finset/NAry.lean index 1ef61280f4ffc..28a21e494d1a0 100644 --- a/Mathlib/Data/Finset/NAry.lean +++ b/Mathlib/Data/Finset/NAry.lean @@ -83,13 +83,19 @@ theorem image_subset_image₂_left (hb : b ∈ t) : s.image (fun a => f a b) ⊆ theorem image_subset_image₂_right (ha : a ∈ s) : t.image (fun b => f a b) ⊆ image₂ f s t := image_subset_iff.2 fun _ => mem_image₂_of_mem ha -theorem forall_image₂_iff {p : γ → Prop} : +lemma forall_mem_image₂ {p : γ → Prop} : (∀ z ∈ image₂ f s t, p z) ↔ ∀ x ∈ s, ∀ y ∈ t, p (f x y) := by - simp_rw [← mem_coe, coe_image₂, forall_image2_iff] + simp_rw [← mem_coe, coe_image₂, forall_mem_image2] + +lemma exists_mem_image₂ {p : γ → Prop} : + (∃ z ∈ image₂ f s t, p z) ↔ ∃ x ∈ s, ∃ y ∈ t, p (f x y) := by + simp_rw [← mem_coe, coe_image₂, exists_mem_image2] + +@[deprecated (since := "2024-11-23")] alias forall_image₂_iff := forall_mem_image₂ @[simp] theorem image₂_subset_iff : image₂ f s t ⊆ u ↔ ∀ x ∈ s, ∀ y ∈ t, f x y ∈ u := - forall_image₂_iff + forall_mem_image₂ theorem image₂_subset_iff_left : image₂ f s t ⊆ u ↔ ∀ a ∈ s, (t.image fun b => f a b) ⊆ u := by simp_rw [image₂_subset_iff, image_subset_iff] @@ -504,10 +510,10 @@ section SemilatticeSup variable [SemilatticeSup δ] -@[simp (default + 1)] -- otherwise `simp` doesn't use `forall_image₂_iff` +@[simp (default + 1)] -- otherwise `simp` doesn't use `forall_mem_image₂` lemma sup'_image₂_le {g : γ → δ} {a : δ} (h : (image₂ f s t).Nonempty) : sup' (image₂ f s t) h g ≤ a ↔ ∀ x ∈ s, ∀ y ∈ t, g (f x y) ≤ a := by - rw [sup'_le_iff, forall_image₂_iff] + rw [sup'_le_iff, forall_mem_image₂] lemma sup'_image₂_left (g : γ → δ) (h : (image₂ f s t).Nonempty) : sup' (image₂ f s t) h g = @@ -521,10 +527,10 @@ lemma sup'_image₂_right (g : γ → δ) (h : (image₂ f s t).Nonempty) : variable [OrderBot δ] -@[simp (default + 1)] -- otherwise `simp` doesn't use `forall_image₂_iff` +@[simp (default + 1)] -- otherwise `simp` doesn't use `forall_mem_image₂` lemma sup_image₂_le {g : γ → δ} {a : δ} : sup (image₂ f s t) g ≤ a ↔ ∀ x ∈ s, ∀ y ∈ t, g (f x y) ≤ a := by - rw [Finset.sup_le_iff, forall_image₂_iff] + rw [Finset.sup_le_iff, forall_mem_image₂] variable (s t) @@ -540,10 +546,10 @@ section SemilatticeInf variable [SemilatticeInf δ] -@[simp (default + 1)] -- otherwise `simp` doesn't use `forall_image₂_iff` +@[simp (default + 1)] -- otherwise `simp` doesn't use `forall_mem_image₂` lemma le_inf'_image₂ {g : γ → δ} {a : δ} (h : (image₂ f s t).Nonempty) : a ≤ inf' (image₂ f s t) h g ↔ ∀ x ∈ s, ∀ y ∈ t, a ≤ g (f x y) := by - rw [le_inf'_iff, forall_image₂_iff] + rw [le_inf'_iff, forall_mem_image₂] lemma inf'_image₂_left (g : γ → δ) (h : (image₂ f s t).Nonempty) : inf' (image₂ f s t) h g = @@ -557,7 +563,7 @@ lemma inf'_image₂_right (g : γ → δ) (h : (image₂ f s t).Nonempty) : variable [OrderTop δ] -@[simp (default + 1)] -- otherwise `simp` doesn't use `forall_image₂_iff` +@[simp (default + 1)] -- otherwise `simp` doesn't use `forall_mem_image₂` lemma le_inf_image₂ {g : γ → δ} {a : δ} : a ≤ inf (image₂ f s t) g ↔ ∀ x ∈ s, ∀ y ∈ t, a ≤ g (f x y) := sup_image₂_le (δ := δᵒᵈ) diff --git a/Mathlib/Data/Finset/Sups.lean b/Mathlib/Data/Finset/Sups.lean index 0b2bf2e8a4359..febbe67418d1c 100644 --- a/Mathlib/Data/Finset/Sups.lean +++ b/Mathlib/Data/Finset/Sups.lean @@ -91,7 +91,7 @@ lemma image_subset_sups_left : b ∈ t → s.image (· ⊔ b) ⊆ s ⊻ t := ima lemma image_subset_sups_right : a ∈ s → t.image (a ⊔ ·) ⊆ s ⊻ t := image_subset_image₂_right theorem forall_sups_iff {p : α → Prop} : (∀ c ∈ s ⊻ t, p c) ↔ ∀ a ∈ s, ∀ b ∈ t, p (a ⊔ b) := - forall_image₂_iff + forall_mem_image₂ @[simp] theorem sups_subset_iff : s ⊻ t ⊆ u ↔ ∀ a ∈ s, ∀ b ∈ t, a ⊔ b ∈ u := @@ -236,7 +236,7 @@ lemma image_subset_infs_left : b ∈ t → s.image (· ⊓ b) ⊆ s ⊼ t := ima lemma image_subset_infs_right : a ∈ s → t.image (a ⊓ ·) ⊆ s ⊼ t := image_subset_image₂_right theorem forall_infs_iff {p : α → Prop} : (∀ c ∈ s ⊼ t, p c) ↔ ∀ a ∈ s, ∀ b ∈ t, p (a ⊓ b) := - forall_image₂_iff + forall_mem_image₂ @[simp] theorem infs_subset_iff : s ⊼ t ⊆ u ↔ ∀ a ∈ s, ∀ b ∈ t, a ⊓ b ∈ u := @@ -547,7 +547,7 @@ lemma image_subset_diffs_left : b ∈ t → s.image (· \ b) ⊆ s \\ t := image lemma image_subset_diffs_right : a ∈ s → t.image (a \ ·) ⊆ s \\ t := image_subset_image₂_right lemma forall_mem_diffs {p : α → Prop} : (∀ c ∈ s \\ t, p c) ↔ ∀ a ∈ s, ∀ b ∈ t, p (a \ b) := - forall_image₂_iff + forall_mem_image₂ @[simp] lemma diffs_subset_iff : s \\ t ⊆ u ↔ ∀ a ∈ s, ∀ b ∈ t, a \ b ∈ u := image₂_subset_iff diff --git a/Mathlib/Data/Set/NAry.lean b/Mathlib/Data/Set/NAry.lean index 72515c986a918..5596fbcdcce5c 100644 --- a/Mathlib/Data/Set/NAry.lean +++ b/Mathlib/Data/Set/NAry.lean @@ -49,13 +49,17 @@ theorem image_subset_image2_left (hb : b ∈ t) : (fun a => f a b) '' s ⊆ imag theorem image_subset_image2_right (ha : a ∈ s) : f a '' t ⊆ image2 f s t := forall_mem_image.2 fun _ => mem_image2_of_mem ha -theorem forall_image2_iff {p : γ → Prop} : - (∀ z ∈ image2 f s t, p z) ↔ ∀ x ∈ s, ∀ y ∈ t, p (f x y) := - ⟨fun h x hx y hy => h _ ⟨x, hx, y, hy, rfl⟩, fun h _ ⟨x, hx, y, hy, hz⟩ => hz ▸ h x hx y hy⟩ +lemma forall_mem_image2 {p : γ → Prop} : + (∀ z ∈ image2 f s t, p z) ↔ ∀ x ∈ s, ∀ y ∈ t, p (f x y) := by aesop + +lemma exists_mem_image2 {p : γ → Prop} : + (∃ z ∈ image2 f s t, p z) ↔ ∃ x ∈ s, ∃ y ∈ t, p (f x y) := by aesop + +@[deprecated (since := "2024-11-23")] alias forall_image2_iff := forall_mem_image2 @[simp] theorem image2_subset_iff {u : Set γ} : image2 f s t ⊆ u ↔ ∀ x ∈ s, ∀ y ∈ t, f x y ∈ u := - forall_image2_iff + forall_mem_image2 theorem image2_subset_iff_left : image2 f s t ⊆ u ↔ ∀ a ∈ s, (fun b => f a b) '' t ⊆ u := by simp_rw [image2_subset_iff, image_subset_iff, subset_def, mem_preimage] @@ -190,7 +194,7 @@ lemma image2_range (f : α' → β' → γ) (g : α → α') (h : β → β') : theorem image2_assoc {f : δ → γ → ε} {g : α → β → δ} {f' : α → ε' → ε} {g' : β → γ → ε'} (h_assoc : ∀ a b c, f (g a b) c = f' a (g' b c)) : image2 f (image2 g s t) u = image2 f' s (image2 g' t u) := - eq_of_forall_subset_iff fun _ ↦ by simp only [image2_subset_iff, forall_image2_iff, h_assoc] + eq_of_forall_subset_iff fun _ ↦ by simp only [image2_subset_iff, forall_mem_image2, h_assoc] theorem image2_comm {g : β → α → γ} (h_comm : ∀ a b, f a b = g b a) : image2 f s t = image2 g t s := (image2_swap _ _ _).trans <| by simp_rw [h_comm] diff --git a/Mathlib/Data/Set/Sups.lean b/Mathlib/Data/Set/Sups.lean index 9c2fe8ef24b07..e9bca4723383e 100644 --- a/Mathlib/Data/Set/Sups.lean +++ b/Mathlib/Data/Set/Sups.lean @@ -90,7 +90,7 @@ theorem image_subset_sups_right : a ∈ s → (· ⊔ ·) a '' t ⊆ s ⊻ t := image_subset_image2_right theorem forall_sups_iff {p : α → Prop} : (∀ c ∈ s ⊻ t, p c) ↔ ∀ a ∈ s, ∀ b ∈ t, p (a ⊔ b) := - forall_image2_iff + forall_mem_image2 @[simp] theorem sups_subset_iff : s ⊻ t ⊆ u ↔ ∀ a ∈ s, ∀ b ∈ t, a ⊔ b ∈ u := @@ -220,7 +220,7 @@ theorem image_subset_infs_right : a ∈ s → (a ⊓ ·) '' t ⊆ s ⊼ t := image_subset_image2_right theorem forall_infs_iff {p : α → Prop} : (∀ c ∈ s ⊼ t, p c) ↔ ∀ a ∈ s, ∀ b ∈ t, p (a ⊓ b) := - forall_image2_iff + forall_mem_image2 @[simp] theorem infs_subset_iff : s ⊼ t ⊆ u ↔ ∀ a ∈ s, ∀ b ∈ t, a ⊓ b ∈ u := diff --git a/Mathlib/Dynamics/OmegaLimit.lean b/Mathlib/Dynamics/OmegaLimit.lean index 7de977bddf6ba..d2ba2df728a80 100644 --- a/Mathlib/Dynamics/OmegaLimit.lean +++ b/Mathlib/Dynamics/OmegaLimit.lean @@ -84,7 +84,7 @@ theorem mapsTo_omegaLimit' {α' β' : Type*} [TopologicalSpace β'] {f : Filter MapsTo gb (ω f ϕ s) (ω f ϕ' s') := by simp only [omegaLimit_def, mem_iInter, MapsTo] intro y hy u hu - refine map_mem_closure hgc (hy _ (inter_mem hu hg)) (forall_image2_iff.2 fun t ht x hx ↦ ?_) + refine map_mem_closure hgc (hy _ (inter_mem hu hg)) (forall_mem_image2.2 fun t ht x hx ↦ ?_) calc ϕ' t (ga x) ∈ image2 ϕ' u s' := mem_image2_of_mem ht.1 (hs hx) _ = gb (ϕ t x) := ht.2 hx |>.symm diff --git a/Mathlib/NumberTheory/MaricaSchoenheim.lean b/Mathlib/NumberTheory/MaricaSchoenheim.lean index d1dbc047e3b2a..bdbe10e46a529 100644 --- a/Mathlib/NumberTheory/MaricaSchoenheim.lean +++ b/Mathlib/NumberTheory/MaricaSchoenheim.lean @@ -47,13 +47,13 @@ lemma grahamConjecture_of_squarefree {n : ℕ} (f : ℕ → ℕ) (hf' : ∀ k < _ < n := tsub_lt_self hn.bot_lt zero_lt_one · rw [Finset.card_image_of_injOn, card_Iio] simpa using prod_primeFactors_invOn_squarefree.2.injOn.comp hf.injOn hf' - · simp only [𝒜, forall_mem_diffs, forall_image, mem_Ioo, mem_Iio] + · simp only [𝒜, forall_mem_diffs, forall_mem_image, mem_Ioo, mem_Iio] rintro i hi j hj rw [← primeFactors_div_gcd (hf' _ hi) (hf' _ hj).ne_zero, prod_primeFactors_of_squarefree <| hf'' _ hi _] exact ⟨Nat.div_pos (gcd_le_left _ (hf' _ hi).ne_zero.bot_lt) <| Nat.gcd_pos_of_pos_left _ (hf' _ hi).ne_zero.bot_lt, Nat.div_lt_of_lt_mul <| this _ hi _ hj⟩ - · simp only [𝒜, Set.InjOn, mem_coe, forall_mem_diffs, forall_image, mem_Ioo, mem_Iio] + · simp only [𝒜, Set.InjOn, mem_coe, forall_mem_diffs, forall_mem_image, mem_Ioo, mem_Iio] rintro a ha b hb c hc d hd rw [← primeFactors_div_gcd (hf' _ ha) (hf' _ hb).ne_zero, ← primeFactors_div_gcd (hf' _ hc) (hf' _ hd).ne_zero, prod_primeFactors_of_squarefree (hf'' _ ha _), diff --git a/Mathlib/Order/Bounds/Image.lean b/Mathlib/Order/Bounds/Image.lean index cb97202b9d3ce..4ef77e7f9e476 100644 --- a/Mathlib/Order/Bounds/Image.lean +++ b/Mathlib/Order/Bounds/Image.lean @@ -248,11 +248,11 @@ include h₀ h₁ theorem mem_upperBounds_image2 (ha : a ∈ upperBounds s) (hb : b ∈ upperBounds t) : f a b ∈ upperBounds (image2 f s t) := - forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy + forall_mem_image2.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy theorem mem_lowerBounds_image2 (ha : a ∈ lowerBounds s) (hb : b ∈ lowerBounds t) : f a b ∈ lowerBounds (image2 f s t) := - forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy + forall_mem_image2.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy theorem image2_upperBounds_upperBounds_subset : image2 f (upperBounds s) (upperBounds t) ⊆ upperBounds (image2 f s t) := @@ -292,11 +292,11 @@ include h₀ h₁ theorem mem_upperBounds_image2_of_mem_upperBounds_of_mem_lowerBounds (ha : a ∈ upperBounds s) (hb : b ∈ lowerBounds t) : f a b ∈ upperBounds (image2 f s t) := - forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy + forall_mem_image2.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy theorem mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_upperBounds (ha : a ∈ lowerBounds s) (hb : b ∈ upperBounds t) : f a b ∈ lowerBounds (image2 f s t) := - forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy + forall_mem_image2.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy theorem image2_upperBounds_lowerBounds_subset_upperBounds_image2 : image2 f (upperBounds s) (lowerBounds t) ⊆ upperBounds (image2 f s t) := @@ -338,11 +338,11 @@ include h₀ h₁ theorem mem_upperBounds_image2_of_mem_lowerBounds (ha : a ∈ lowerBounds s) (hb : b ∈ lowerBounds t) : f a b ∈ upperBounds (image2 f s t) := - forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy + forall_mem_image2.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy theorem mem_lowerBounds_image2_of_mem_upperBounds (ha : a ∈ upperBounds s) (hb : b ∈ upperBounds t) : f a b ∈ lowerBounds (image2 f s t) := - forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy + forall_mem_image2.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy theorem image2_upperBounds_upperBounds_subset_upperBounds_image2 : image2 f (lowerBounds s) (lowerBounds t) ⊆ upperBounds (image2 f s t) := @@ -380,11 +380,11 @@ include h₀ h₁ theorem mem_upperBounds_image2_of_mem_upperBounds_of_mem_upperBounds (ha : a ∈ lowerBounds s) (hb : b ∈ upperBounds t) : f a b ∈ upperBounds (image2 f s t) := - forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy + forall_mem_image2.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy theorem mem_lowerBounds_image2_of_mem_lowerBounds_of_mem_lowerBounds (ha : a ∈ upperBounds s) (hb : b ∈ lowerBounds t) : f a b ∈ lowerBounds (image2 f s t) := - forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy + forall_mem_image2.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy theorem image2_lowerBounds_upperBounds_subset_upperBounds_image2 : image2 f (lowerBounds s) (upperBounds t) ⊆ upperBounds (image2 f s t) := diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean index 437a1e3d52e13..0773cfe1e11e4 100644 --- a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean +++ b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean @@ -809,7 +809,7 @@ theorem csSup_image2_eq_csSup_csSup (h₁ : ∀ b, GaloisConnection (swap l b) ( refine eq_of_forall_ge_iff fun c => ?_ rw [csSup_le_iff (hs₁.image2 (fun _ => (h₁ _).monotone_l) (fun _ => (h₂ _).monotone_l) ht₁) (hs₀.image2 ht₀), - forall_image2_iff, forall₂_swap, (h₂ _).le_iff_le, csSup_le_iff ht₁ ht₀] + forall_mem_image2, forall₂_swap, (h₂ _).le_iff_le, csSup_le_iff ht₁ ht₀] simp_rw [← (h₂ _).le_iff_le, (h₁ _).le_iff_le, csSup_le_iff hs₁ hs₀] theorem csSup_image2_eq_csSup_csInf (h₁ : ∀ b, GaloisConnection (swap l b) (u₁ b)) diff --git a/Mathlib/Topology/CompactOpen.lean b/Mathlib/Topology/CompactOpen.lean index f3479618c96ad..49f2b425934f9 100644 --- a/Mathlib/Topology/CompactOpen.lean +++ b/Mathlib/Topology/CompactOpen.lean @@ -72,7 +72,7 @@ lemma tendsto_nhds_compactOpen {l : Filter α} {f : α → C(Y, Z)} {g : C(Y, Z) lemma continuous_compactOpen {f : X → C(Y, Z)} : Continuous f ↔ ∀ K, IsCompact K → ∀ U, IsOpen U → IsOpen {x | MapsTo (f x) K U} := - continuous_generateFrom_iff.trans forall_image2_iff + continuous_generateFrom_iff.trans forall_mem_image2 section Functorial @@ -274,7 +274,7 @@ theorem compactOpen_eq_iInf_induced : (ContinuousMap.compactOpen : TopologicalSpace C(X, Y)) = ⨅ (K : Set X) (_ : IsCompact K), .induced (.restrict K) ContinuousMap.compactOpen := by refine le_antisymm (le_iInf₂ fun s _ ↦ compactOpen_le_induced s) ?_ - refine le_generateFrom <| forall_image2_iff.2 fun K (hK : IsCompact K) U hU ↦ ?_ + refine le_generateFrom <| forall_mem_image2.2 fun K (hK : IsCompact K) U hU ↦ ?_ refine TopologicalSpace.le_def.1 (iInf₂_le K hK) _ ?_ convert isOpen_induced (isOpen_setOf_mapsTo (isCompact_iff_isCompact_univ.1 hK) hU) simp [mapsTo_univ_iff, Subtype.forall, MapsTo] diff --git a/Mathlib/Topology/EMetricSpace/Lipschitz.lean b/Mathlib/Topology/EMetricSpace/Lipschitz.lean index a9c56f4bc73d9..3be6db91a4e89 100644 --- a/Mathlib/Topology/EMetricSpace/Lipschitz.lean +++ b/Mathlib/Topology/EMetricSpace/Lipschitz.lean @@ -320,7 +320,7 @@ protected theorem prod {g : α → γ} {Kf Kg : ℝ≥0} (hf : LipschitzOnWith K theorem ediam_image2_le (f : α → β → γ) {K₁ K₂ : ℝ≥0} (s : Set α) (t : Set β) (hf₁ : ∀ b ∈ t, LipschitzOnWith K₁ (f · b) s) (hf₂ : ∀ a ∈ s, LipschitzOnWith K₂ (f a) t) : EMetric.diam (Set.image2 f s t) ≤ ↑K₁ * EMetric.diam s + ↑K₂ * EMetric.diam t := by - simp only [EMetric.diam_le_iff, forall_image2_iff] + simp only [EMetric.diam_le_iff, forall_mem_image2] intro a₁ ha₁ b₁ hb₁ a₂ ha₂ b₂ hb₂ refine (edist_triangle _ (f a₂ b₁) _).trans ?_ exact From 14375b940a8d2372f6eff562acf72b9a75393810 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Tue, 10 Dec 2024 15:20:23 +0000 Subject: [PATCH 636/829] refactor/fix: don't use RingHom.toAlgebra in Ideal.quotientAlgebra (#19847) --- .../RamificationInertia/Basic.lean | 113 +++++++++--------- Mathlib/RingTheory/Ideal/Over.lean | 12 +- .../RingTheory/Ideal/Quotient/Operations.lean | 14 ++- 3 files changed, 77 insertions(+), 62 deletions(-) diff --git a/Mathlib/NumberTheory/RamificationInertia/Basic.lean b/Mathlib/NumberTheory/RamificationInertia/Basic.lean index 92cb53f8ad16f..d0291b1916d6f 100644 --- a/Mathlib/NumberTheory/RamificationInertia/Basic.lean +++ b/Mathlib/NumberTheory/RamificationInertia/Basic.lean @@ -181,7 +181,9 @@ theorem ramificationIdx_ne_zero (hp0 : map f p ≠ ⊥) (hP : P.IsPrime) (le : m end IsDedekindDomain -variable (f p P) +variable (f p P) [Algebra R S] + +local notation "f" => algebraMap R S open Classical in /-- The inertia degree of `P : Ideal S` lying over `p : Ideal R` is the degree of the @@ -194,31 +196,31 @@ and there is an algebra structure `R / p → S / P`. -/ noncomputable def inertiaDeg : ℕ := if hPp : comap f P = p then - letI : Algebra (R ⧸ p) (S ⧸ P) := Quotient.algebraQuotientOfLEComap (le_of_eq hPp.symm) + letI : Algebra (R ⧸ p) (S ⧸ P) := Quotient.algebraQuotientOfLEComap hPp.ge finrank (R ⧸ p) (S ⧸ P) else 0 -- Useful for the `nontriviality` tactic using `comap_eq_of_scalar_tower_quotient`. @[simp] theorem inertiaDeg_of_subsingleton [hp : p.IsMaximal] [hQ : Subsingleton (S ⧸ P)] : - inertiaDeg f p P = 0 := by + inertiaDeg p P = 0 := by have := Ideal.Quotient.subsingleton_iff.mp hQ subst this exact dif_neg fun h => hp.ne_top <| h.symm.trans comap_top @[simp] -theorem inertiaDeg_algebraMap [Algebra R S] [P.LiesOver p] [p.IsMaximal] : - inertiaDeg (algebraMap R S) p P = finrank (R ⧸ p) (S ⧸ P) := by +theorem inertiaDeg_algebraMap [P.LiesOver p] [p.IsMaximal] : + inertiaDeg p P = finrank (R ⧸ p) (S ⧸ P) := by nontriviality S ⧸ P using inertiaDeg_of_subsingleton, finrank_zero_of_subsingleton rw [inertiaDeg, dif_pos (over_def P p).symm] -theorem inertiaDeg_pos [p.IsMaximal] [Algebra R S] [Module.Finite R S] - [P.LiesOver p] : 0 < inertiaDeg (algebraMap R S) p P := +theorem inertiaDeg_pos [p.IsMaximal] [Module.Finite R S] + [P.LiesOver p] : 0 < inertiaDeg p P := haveI : Nontrivial (S ⧸ P) := Quotient.nontrivial_of_liesOver_of_isPrime P p finrank_pos.trans_eq (inertiaDeg_algebraMap p P).symm -lemma inertiaDeg_comap_eq [Algebra R S] (e : S ≃ₐ[R] S₁) (P : Ideal S₁) [p.IsMaximal] : - inertiaDeg (algebraMap R S) p (P.comap e) = inertiaDeg (algebraMap R S₁) p P := by +lemma inertiaDeg_comap_eq (e : S ≃ₐ[R] S₁) (P : Ideal S₁) [p.IsMaximal] : + inertiaDeg p (P.comap e) = inertiaDeg p P := by have he : (P.comap e).comap (algebraMap R S) = p ↔ P.comap (algebraMap R S₁) = p := by rw [← comap_coe e, comap_comap, ← e.toAlgHom_toRingHom, AlgHom.comp_algebraMap] by_cases h : P.LiesOver p @@ -227,9 +229,9 @@ lemma inertiaDeg_comap_eq [Algebra R S] (e : S ≃ₐ[R] S₁) (P : Ideal S₁) · rw [inertiaDeg, dif_neg (fun eq => h ⟨(he.mp eq).symm⟩)] rw [inertiaDeg, dif_neg (fun eq => h ⟨eq.symm⟩)] -lemma inertiaDeg_map_eq [p.IsMaximal] [Algebra R S] (P : Ideal S) +lemma inertiaDeg_map_eq [p.IsMaximal] (P : Ideal S) {E : Type*} [EquivLike E S S₁] [AlgEquivClass E R S S₁] (e : E) : - inertiaDeg (algebraMap R S₁) p (P.map e) = inertiaDeg (algebraMap R S) p P := by + inertiaDeg p (P.map e) = inertiaDeg p P := by rw [show P.map e = _ from map_comap_of_equiv (e : S ≃+* S₁)] exact p.inertiaDeg_comap_eq (e : S ≃ₐ[R] S₁).symm P @@ -447,6 +449,9 @@ end FinrankQuotientMap section FactLeComap +variable [Algebra R S] + +local notation "f" => algebraMap R S local notation "e" => ramificationIdx f p P /-- `R / p` has a canonical map to `S / (P ^ e)`, where `e` is the ramification index @@ -462,17 +467,15 @@ theorem Quotient.algebraMap_quotient_pow_ramificationIdx (x : R) : This can't be an instance since the map `f : R → S` is generally not inferrable. -/ -def Quotient.algebraQuotientOfRamificationIdxNeZero [hfp : NeZero (ramificationIdx f p P)] : +def Quotient.algebraQuotientOfRamificationIdxNeZero [hfp : NeZero e] : Algebra (R ⧸ p) (S ⧸ P) := Quotient.algebraQuotientOfLEComap (le_comap_of_ramificationIdx_ne_zero hfp.out) -set_option synthInstance.checkSynthOrder false -- Porting note: this is okay by the remark below --- In this file, the value for `f` can be inferred. attribute [local instance] Ideal.Quotient.algebraQuotientOfRamificationIdxNeZero @[simp] theorem Quotient.algebraMap_quotient_of_ramificationIdx_neZero - [NeZero (ramificationIdx f p P)] (x : R) : + [NeZero e] (x : R) : algebraMap (R ⧸ p) (S ⧸ P) (Ideal.Quotient.mk p x) = Ideal.Quotient.mk P (f x) := rfl /-- The inclusion `(P^(i + 1) / P^e) ⊂ (P^i / P^e)`. -/ @@ -485,7 +488,7 @@ def powQuotSuccInclusion (i : ℕ) : map_smul' _ _ := rfl theorem powQuotSuccInclusion_injective (i : ℕ) : - Function.Injective (powQuotSuccInclusion f p P i) := by + Function.Injective (powQuotSuccInclusion p P i) := by rw [← LinearMap.ker_eq_bot, LinearMap.ker_eq_bot'] rintro ⟨x, hx⟩ hx0 rw [Subtype.ext_iff] at hx0 ⊢ @@ -497,7 +500,7 @@ and `quotientRangePowQuotSuccInclusionEquiv` for this as a linear equivalence. -/ noncomputable def quotientToQuotientRangePowQuotSuccAux {i : ℕ} {a : S} (a_mem : a ∈ P ^ i) : S ⧸ P → - (P ^ i).map (Ideal.Quotient.mk (P ^ e)) ⧸ LinearMap.range (powQuotSuccInclusion f p P i) := + (P ^ i).map (Ideal.Quotient.mk (P ^ e)) ⧸ LinearMap.range (powQuotSuccInclusion p P i) := Quotient.map' (fun x : S => ⟨_, Ideal.mem_map_of_mem _ (Ideal.mul_mem_right x _ a_mem)⟩) fun x y h => by rw [Submodule.quotientRel_def] at h ⊢ @@ -508,18 +511,19 @@ noncomputable def quotientToQuotientRangePowQuotSuccAux {i : ℕ} {a : S} (a_mem Subtype.coe_mk, _root_.map_mul, map_sub, mul_sub] theorem quotientToQuotientRangePowQuotSuccAux_mk {i : ℕ} {a : S} (a_mem : a ∈ P ^ i) (x : S) : - quotientToQuotientRangePowQuotSuccAux f p P a_mem (Submodule.Quotient.mk x) = + quotientToQuotientRangePowQuotSuccAux p P a_mem (Submodule.Quotient.mk x) = Submodule.Quotient.mk ⟨_, Ideal.mem_map_of_mem _ (Ideal.mul_mem_right x _ a_mem)⟩ := by apply Quotient.map'_mk'' section -variable [hfp : NeZero (ramificationIdx f p P)] +variable [hfp : NeZero (ramificationIdx (algebraMap R S) p P)] /-- `S ⧸ P` embeds into the quotient by `P^(i+1) ⧸ P^e` as a subspace of `P^i ⧸ P^e`. -/ -noncomputable def quotientToQuotientRangePowQuotSucc {i : ℕ} {a : S} (a_mem : a ∈ P ^ i) : +noncomputable def quotientToQuotientRangePowQuotSucc + {i : ℕ} {a : S} (a_mem : a ∈ P ^ i) : S ⧸ P →ₗ[R ⧸ p] - (P ^ i).map (Ideal.Quotient.mk (P ^ e)) ⧸ LinearMap.range (powQuotSuccInclusion f p P i) where - toFun := quotientToQuotientRangePowQuotSuccAux f p P a_mem + (P ^ i).map (Ideal.Quotient.mk (P ^ e)) ⧸ LinearMap.range (powQuotSuccInclusion p P i) where + toFun := quotientToQuotientRangePowQuotSuccAux p P a_mem map_add' := by intro x y; refine Quotient.inductionOn' x fun x => Quotient.inductionOn' y fun y => ?_ simp only [Submodule.Quotient.mk''_eq_mk, ← Submodule.Quotient.mk_add, @@ -536,13 +540,13 @@ noncomputable def quotientToQuotientRangePowQuotSucc {i : ℕ} {a : S} (a_mem : ring theorem quotientToQuotientRangePowQuotSucc_mk {i : ℕ} {a : S} (a_mem : a ∈ P ^ i) (x : S) : - quotientToQuotientRangePowQuotSucc f p P a_mem (Submodule.Quotient.mk x) = + quotientToQuotientRangePowQuotSucc p P a_mem (Submodule.Quotient.mk x) = Submodule.Quotient.mk ⟨_, Ideal.mem_map_of_mem _ (Ideal.mul_mem_right x _ a_mem)⟩ := - quotientToQuotientRangePowQuotSuccAux_mk f p P a_mem x + quotientToQuotientRangePowQuotSuccAux_mk p P a_mem x theorem quotientToQuotientRangePowQuotSucc_injective [IsDedekindDomain S] [P.IsPrime] {i : ℕ} (hi : i < e) {a : S} (a_mem : a ∈ P ^ i) (a_not_mem : a ∉ P ^ (i + 1)) : - Function.Injective (quotientToQuotientRangePowQuotSucc f p P a_mem) := fun x => + Function.Injective (quotientToQuotientRangePowQuotSucc p P a_mem) := fun x => Quotient.inductionOn' x fun x y => Quotient.inductionOn' y fun y h => by have Pe_le_Pi1 : P ^ e ≤ P ^ (i + 1) := Ideal.pow_le_pow_right hi @@ -562,7 +566,7 @@ theorem quotientToQuotientRangePowQuotSucc_injective [IsDedekindDomain S] [P.IsP theorem quotientToQuotientRangePowQuotSucc_surjective [IsDedekindDomain S] (hP0 : P ≠ ⊥) [hP : P.IsPrime] {i : ℕ} (hi : i < e) {a : S} (a_mem : a ∈ P ^ i) (a_not_mem : a ∉ P ^ (i + 1)) : - Function.Surjective (quotientToQuotientRangePowQuotSucc f p P a_mem) := by + Function.Surjective (quotientToQuotientRangePowQuotSucc p P a_mem) := by rintro ⟨⟨⟨x⟩, hx⟩⟩ have Pe_le_Pi : P ^ e ≤ P ^ i := Ideal.pow_le_pow_right hi.le rw [Submodule.Quotient.quot_mk_eq_mk, Ideal.Quotient.mk_eq_mk, Ideal.mem_quotient_iff_mem_sup, @@ -592,15 +596,15 @@ theorem quotientToQuotientRangePowQuotSucc_surjective [IsDedekindDomain S] `R ⧸ p`-linearly isomorphic to `S ⧸ P`. -/ noncomputable def quotientRangePowQuotSuccInclusionEquiv [IsDedekindDomain S] [P.IsPrime] (hP : P ≠ ⊥) {i : ℕ} (hi : i < e) : - ((P ^ i).map (Ideal.Quotient.mk (P ^ e)) ⧸ LinearMap.range (powQuotSuccInclusion f p P i)) + ((P ^ i).map (Ideal.Quotient.mk (P ^ e)) ⧸ LinearMap.range (powQuotSuccInclusion p P i)) ≃ₗ[R ⧸ p] S ⧸ P := by choose a a_mem a_not_mem using SetLike.exists_of_lt (Ideal.pow_right_strictAnti P hP (Ideal.IsPrime.ne_top inferInstance) (le_refl i.succ)) refine (LinearEquiv.ofBijective ?_ ⟨?_, ?_⟩).symm - · exact quotientToQuotientRangePowQuotSucc f p P a_mem - · exact quotientToQuotientRangePowQuotSucc_injective f p P hi a_mem a_not_mem - · exact quotientToQuotientRangePowQuotSucc_surjective f p P hP hi a_mem a_not_mem + · exact quotientToQuotientRangePowQuotSucc p P a_mem + · exact quotientToQuotientRangePowQuotSucc_injective p P hi a_mem a_not_mem + · exact quotientToQuotientRangePowQuotSucc_surjective p P hP hi a_mem a_not_mem /-- Since the inclusion `(P^(i + 1) / P^e) ⊂ (P^i / P^e)` has a kernel isomorphic to `P / S`, `[P^i / P^e : R / p] = [P^(i+1) / P^e : R / p] + [P / S : R / p]` -/ @@ -609,9 +613,9 @@ theorem rank_pow_quot_aux [IsDedekindDomain S] [p.IsMaximal] [P.IsPrime] (hP0 : Module.rank (R ⧸ p) (Ideal.map (Ideal.Quotient.mk (P ^ e)) (P ^ i)) = Module.rank (R ⧸ p) (S ⧸ P) + Module.rank (R ⧸ p) (Ideal.map (Ideal.Quotient.mk (P ^ e)) (P ^ (i + 1))) := by - rw [← rank_range_of_injective _ (powQuotSuccInclusion_injective f p P i), - (quotientRangePowQuotSuccInclusionEquiv f p P hP0 hi).symm.rank_eq] - exact (Submodule.rank_quotient_add_rank (LinearMap.range (powQuotSuccInclusion f p P i))).symm + rw [← rank_range_of_injective _ (powQuotSuccInclusion_injective p P i), + (quotientRangePowQuotSuccInclusionEquiv p P hP0 hi).symm.rank_eq] + exact (Submodule.rank_quotient_add_rank (LinearMap.range (powQuotSuccInclusion p P i))).symm theorem rank_pow_quot [IsDedekindDomain S] [p.IsMaximal] [P.IsPrime] (hP0 : P ≠ ⊥) (i : ℕ) (hi : i ≤ e) : @@ -623,7 +627,7 @@ theorem rank_pow_quot [IsDedekindDomain S] [p.IsMaximal] [P.IsPrime] (hP0 : P = (e - i) • Module.rank (R ⧸ p) (S ⧸ P) refine Nat.decreasingInduction' (P := Q) (fun j lt_e _le_j ih => ?_) hi ?_ · dsimp only [Q] - rw [rank_pow_quot_aux f p P _ lt_e, ih, ← succ_nsmul', Nat.sub_succ, ← Nat.succ_eq_add_one, + rw [rank_pow_quot_aux p P _ lt_e, ih, ← succ_nsmul', Nat.sub_succ, ← Nat.succ_eq_add_one, Nat.succ_pred_eq_of_pos (Nat.sub_pos_of_lt lt_e)] assumption · dsimp only [Q] @@ -642,7 +646,7 @@ theorem rank_prime_pow_ramificationIdx [IsDedekindDomain S] [p.IsMaximal] [P.IsP (@Algebra.toModule _ _ _ _ <| @Quotient.algebraQuotientOfRamificationIdxNeZero _ _ _ _ _ _ _ ⟨he⟩) := by letI : NeZero e := ⟨he⟩ - have := rank_pow_quot f p P hP0 0 (Nat.zero_le e) + have := rank_pow_quot p P hP0 0 (Nat.zero_le e) rw [pow_zero, Nat.sub_zero, Ideal.one_eq_top, Ideal.map_top] at this exact (rank_top (R ⧸ p) _).symm.trans this @@ -656,8 +660,8 @@ theorem finrank_prime_pow_ramificationIdx [IsDedekindDomain S] (hP0 : P ≠ ⊥) (@Algebra.toModule _ _ _ _ <| @Quotient.algebraQuotientOfRamificationIdxNeZero _ _ _ _ _ _ _ ⟨he⟩) := by letI : NeZero e := ⟨he⟩ - letI : Algebra (R ⧸ p) (S ⧸ P) := Quotient.algebraQuotientOfRamificationIdxNeZero f p P - have hdim := rank_prime_pow_ramificationIdx _ _ _ hP0 he + letI : Algebra (R ⧸ p) (S ⧸ P) := Quotient.algebraQuotientOfRamificationIdxNeZero p P + have hdim := rank_prime_pow_ramificationIdx _ _ hP0 he by_cases hP : FiniteDimensional (R ⧸ p) (S ⧸ P) · haveI := hP haveI := (finiteDimensional_iff_of_rank_eq_nsmul he hdim).mpr hP @@ -709,7 +713,7 @@ instance Factors.liesOver [p.IsMaximal] (P : (factors (map (algebraMap R S) p)). theorem Factors.finrank_pow_ramificationIdx [p.IsMaximal] (P : (factors (map (algebraMap R S) p)).toFinset) : finrank (R ⧸ p) (S ⧸ (P : Ideal S) ^ ramificationIdx (algebraMap R S) p P) = - ramificationIdx (algebraMap R S) p P * inertiaDeg (algebraMap R S) p P := by + ramificationIdx (algebraMap R S) p P * inertiaDeg p (P : Ideal S) := by rw [finrank_prime_pow_ramificationIdx, inertiaDeg_algebraMap] exacts [Factors.ne_bot p P, NeZero.ne _] @@ -772,10 +776,10 @@ theorem sum_ramification_inertia (K L : Type*) [Field K] [Field L] [IsDedekindDo [Algebra R L] [IsScalarTower R S L] [IsScalarTower R K L] [Module.Finite R S] [p.IsMaximal] (hp0 : p ≠ ⊥) : (∑ P ∈ (factors (map (algebraMap R S) p)).toFinset, - ramificationIdx (algebraMap R S) p P * inertiaDeg (algebraMap R S) p P) = + ramificationIdx (algebraMap R S) p P * inertiaDeg p P) = finrank K L := by set e := ramificationIdx (algebraMap R S) p - set f := inertiaDeg (algebraMap R S) p + set f := inertiaDeg p (S := S) calc (∑ P ∈ (factors (map (algebraMap R S) p)).toFinset, e P * f P) = ∑ P ∈ (factors (map (algebraMap R S) p)).toFinset.attach, @@ -822,17 +826,6 @@ theorem ramificationIdx_tower [IsDedekindDomain S] [IsDedekindDomain T] {f : R exact add_right_eq_self.mpr <| Decidable.byContradiction fun h ↦ hntq <| hcp.trans_le <| sup_le hg <| le_of_dvd <| dvd_of_mem_normalizedFactors <| Multiset.count_ne_zero.mp h -theorem inertiaDeg_tower {f : R →+* S} {g : S →+* T} {p : Ideal R} {P : Ideal S} {I : Ideal T} - [p.IsMaximal] [P.IsMaximal] (hp : p = comap f P) (hP : P = comap g I) : - inertiaDeg (g.comp f) p I = inertiaDeg f p P * inertiaDeg g P I := by - have h : comap (g.comp f) I = p := by rw [hp, hP, comap_comap] - simp only [inertiaDeg, dif_pos hp.symm, dif_pos hP.symm, dif_pos h] - letI : Algebra (R ⧸ p) (S ⧸ P) := Ideal.Quotient.algebraQuotientOfLEComap (le_of_eq hp) - letI : Algebra (S ⧸ P) (T ⧸ I) := Ideal.Quotient.algebraQuotientOfLEComap (le_of_eq hP) - letI : Algebra (R ⧸ p) (T ⧸ I) := Ideal.Quotient.algebraQuotientOfLEComap (le_of_eq h.symm) - letI : IsScalarTower (R ⧸ p) (S ⧸ P) (T ⧸ I) := IsScalarTower.of_algebraMap_eq (by rintro ⟨⟩; rfl) - exact (finrank_mul_finrank (R ⧸ p) (S ⧸ P) (T ⧸ I)).symm - variable [Algebra R S] [Algebra S T] [Algebra R T] [IsScalarTower R S T] /-- Let `T / S / R` be a tower of algebras, `p, P, Q` be ideals in `R, S, T` respectively, @@ -850,10 +843,20 @@ theorem ramificationIdx_algebra_tower [IsDedekindDomain S] [IsDedekindDomain T] and `p` and `P` are maximal. If `p = P ∩ S` and `P = I ∩ S`, then `f (I | p) = f (P | p) * f (I | P)`. -/ theorem inertiaDeg_algebra_tower (p : Ideal R) (P : Ideal S) (I : Ideal T) [p.IsMaximal] - [P.IsMaximal] [P.LiesOver p] [I.LiesOver P] : inertiaDeg (algebraMap R T) p I = - inertiaDeg (algebraMap R S) p P * inertiaDeg (algebraMap S T) P I := by - rw [IsScalarTower.algebraMap_eq R S T] - exact inertiaDeg_tower (over_def P p) (over_def I P) + [P.IsMaximal] [P.LiesOver p] [I.LiesOver P] : inertiaDeg p I = + inertiaDeg p P * inertiaDeg P I := by + have h₁ := P.over_def p + have h₂ := I.over_def P + have h₃ := (LiesOver.trans I P p).over + simp only [inertiaDeg, dif_pos h₁.symm, dif_pos h₂.symm, dif_pos h₃.symm] + letI : Algebra (R ⧸ p) (S ⧸ P) := Ideal.Quotient.algebraQuotientOfLEComap h₁.le + letI : Algebra (S ⧸ P) (T ⧸ I) := Ideal.Quotient.algebraQuotientOfLEComap h₂.le + letI : Algebra (R ⧸ p) (T ⧸ I) := Ideal.Quotient.algebraQuotientOfLEComap h₃.le + letI : IsScalarTower (R ⧸ p) (S ⧸ P) (T ⧸ I) := IsScalarTower.of_algebraMap_eq <| by + rintro ⟨x⟩; exact congr_arg _ (IsScalarTower.algebraMap_apply R S T x) + exact (finrank_mul_finrank (R ⧸ p) (S ⧸ P) (T ⧸ I)).symm + +@[deprecated (since := "2024-12-09")] alias inertiaDeg_tower := inertiaDeg_algebra_tower end tower diff --git a/Mathlib/RingTheory/Ideal/Over.lean b/Mathlib/RingTheory/Ideal/Over.lean index 3442658846e0f..876ab2aa23179 100644 --- a/Mathlib/RingTheory/Ideal/Over.lean +++ b/Mathlib/RingTheory/Ideal/Over.lean @@ -138,24 +138,24 @@ theorem comap_eq_of_scalar_tower_quotient [Algebra R S] [Algebra (R ⧸ p) (S · intro hx rw [hx, RingHom.map_zero] -/-- If `P` lies over `p`, then `R / p` has a canonical map to `S / P`. -/ -def Quotient.algebraQuotientOfLEComap (h : p ≤ comap f P) : Algebra (R ⧸ p) (S ⧸ P) := - RingHom.toAlgebra <| quotientMap _ f h +variable [Algebra R S] /-- `R / p` has a canonical map to `S / pS`. -/ -instance Quotient.algebraQuotientMapQuotient : Algebra (R ⧸ p) (S ⧸ map f p) := +instance Quotient.algebraQuotientMapQuotient : Algebra (R ⧸ p) (S ⧸ map (algebraMap R S) p) := Ideal.Quotient.algebraQuotientOfLEComap le_comap_map @[simp] theorem Quotient.algebraMap_quotient_map_quotient (x : R) : + letI f := algebraMap R S algebraMap (R ⧸ p) (S ⧸ map f p) (Ideal.Quotient.mk p x) = Ideal.Quotient.mk (map f p) (f x) := rfl @[simp] theorem Quotient.mk_smul_mk_quotient_map_quotient (x : R) (y : S) : + letI f := algebraMap R S Quotient.mk p x • Quotient.mk (map f p) y = Quotient.mk (map f p) (f x * y) := - rfl + Algebra.smul_def _ _ instance Quotient.tower_quotient_map_quotient [Algebra R S] : IsScalarTower R (R ⧸ p) (S ⧸ map (algebraMap R S) p) := @@ -164,7 +164,7 @@ instance Quotient.tower_quotient_map_quotient [Algebra R S] : Quotient.mk_algebraMap] instance QuotientMapQuotient.isNoetherian [Algebra R S] [IsNoetherian R S] (I : Ideal R) : - IsNoetherian (R ⧸ I) (S ⧸ Ideal.map (algebraMap R S) I) := + IsNoetherian (R ⧸ I) (S ⧸ I.map (algebraMap R S)) := isNoetherian_of_tower R <| isNoetherian_of_surjective S (Ideal.Quotient.mkₐ R _).toLinearMap <| LinearMap.range_eq_top.mpr Ideal.Quotient.mk_surjective diff --git a/Mathlib/RingTheory/Ideal/Quotient/Operations.lean b/Mathlib/RingTheory/Ideal/Quotient/Operations.lean index 8d8291b64c232..488ffed5743b2 100644 --- a/Mathlib/RingTheory/Ideal/Quotient/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Quotient/Operations.lean @@ -590,9 +590,21 @@ def quotientEquivAlg (I : Ideal A) (J : Ideal B) (f : A ≃ₐ[R₁] B) (hIJ : J end +/-- If `P` lies over `p`, then `R / p` has a canonical map to `A / P`. -/ +abbrev Quotient.algebraQuotientOfLEComap [Algebra R A] {p : Ideal R} {P : Ideal A} + (h : p ≤ comap (algebraMap R A) P) : Algebra (R ⧸ p) (A ⧸ P) where + toRingHom := quotientMap P (algebraMap R A) h + smul := Quotient.lift₂ (⟦· • ·⟧) fun r₁ a₁ r₂ a₂ hr ha ↦ Quotient.sound <| by + have := h (p.quotientRel_def.mp hr) + rw [mem_comap, map_sub] at this + simpa only [Algebra.smul_def] using P.quotientRel_def.mpr + (P.mul_sub_mul_mem this <| P.quotientRel_def.mp ha) + smul_def' := by rintro ⟨_⟩ ⟨_⟩; exact congr_arg (⟦·⟧) (Algebra.smul_def _ _) + commutes' := by rintro ⟨_⟩ ⟨_⟩; exact congr_arg (⟦·⟧) (Algebra.commutes _ _) + instance (priority := 100) quotientAlgebra {I : Ideal A} [Algebra R A] : Algebra (R ⧸ I.comap (algebraMap R A)) (A ⧸ I) := - (quotientMap I (algebraMap R A) (le_of_eq rfl)).toAlgebra + Quotient.algebraQuotientOfLEComap le_rfl theorem algebraMap_quotient_injective {I : Ideal A} [Algebra R A] : Function.Injective (algebraMap (R ⧸ I.comap (algebraMap R A)) (A ⧸ I)) := by From afe7fb4d22071d8964f93c65f8b9885be071dee4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 10 Dec 2024 15:41:13 +0000 Subject: [PATCH 637/829] =?UTF-8?q?feat:=20for=20`s`=20a=20subgroup,=20`0?= =?UTF-8?q?=20=E2=80=A2=20s=20=3D=20=E2=8A=A5`=20(#19866)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From FLT --- Mathlib/Algebra/Group/Subgroup/Pointwise.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean index b5e3e4461137b..59e702a1824e6 100644 --- a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean +++ b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean @@ -595,6 +595,14 @@ theorem le_pointwise_smul_iff₀ {a : α} (ha : a ≠ 0) {S T : AddSubgroup A} : end GroupWithZero +section Semiring +variable {R M : Type*} [Semiring R] [AddCommGroup M] [Module R M] + +@[simp] protected lemma zero_smul (s : AddSubgroup M) : (0 : R) • s = ⊥ := by + simp [eq_bot_iff_forall, pointwise_smul_def] + +end Semiring + section Mul variable {R : Type*} [NonUnitalNonAssocRing R] From d70fcf43625e913424f156c19d1f378ad5117f02 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 10 Dec 2024 15:41:14 +0000 Subject: [PATCH 638/829] feat: `Units.map` of `Units.mk` (#19867) From FLT --- Mathlib/Algebra/Group/Units/Hom.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Mathlib/Algebra/Group/Units/Hom.lean b/Mathlib/Algebra/Group/Units/Hom.lean index f82c97cddb10c..a2c7a435b6886 100644 --- a/Mathlib/Algebra/Group/Units/Hom.lean +++ b/Mathlib/Algebra/Group/Units/Hom.lean @@ -79,6 +79,11 @@ theorem coe_map (f : M →* N) (x : Mˣ) : ↑(map f x) = f x := rfl @[to_additive (attr := simp)] theorem coe_map_inv (f : M →* N) (u : Mˣ) : ↑(map f u)⁻¹ = f ↑u⁻¹ := rfl +@[to_additive (attr := simp)] +lemma map_mk (f : M →* N) (val inv : M) (val_inv inv_val) : + map f (mk val inv val_inv inv_val) = mk (f val) (f inv) + (by rw [← f.map_mul, val_inv, f.map_one]) (by rw [← f.map_mul, inv_val, f.map_one]) := rfl + @[to_additive (attr := simp)] theorem map_comp (f : M →* N) (g : N →* P) : map (g.comp f) = (map g).comp (map f) := rfl From b83e52e0c3f96dcf9962bb68198371e761809018 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 10 Dec 2024 16:18:08 +0000 Subject: [PATCH 639/829] =?UTF-8?q?feat(SetTheory/Ordinal/Arithmetic):=20`?= =?UTF-8?q?a=20+=20b=20=E2=89=A4=20c=20=E2=86=94=20=E2=88=80=20d=20<=20b,?= =?UTF-8?q?=20a=20+=20d=20<=20c`=20(#19804)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 23 ++++++++++------------- 1 file changed, 10 insertions(+), 13 deletions(-) diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index b01376da2f420..4a4f0e4cbacc1 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -468,7 +468,8 @@ theorem IsNormal.isLimit {f} (H : IsNormal f) {o} (ho : IsLimit o) : IsLimit (f apply hab.trans_lt rwa [H.lt_iff] -theorem add_le_of_limit {a b c : Ordinal} (h : IsLimit b) : a + b ≤ c ↔ ∀ b' < b, a + b' ≤ c := +private theorem add_le_of_limit {a b c : Ordinal} (h : IsLimit b) : + a + b ≤ c ↔ ∀ b' < b, a + b' ≤ c := ⟨fun h _ l => (add_le_add_left l.le _).trans h, fun H => le_of_not_lt <| by -- Porting note: `induction` tactics are required because of the parser bug. @@ -593,6 +594,14 @@ theorem lt_add_iff {a b c : Ordinal} (hc : c ≠ 0) : a < b + c ↔ ∃ d < c, a rintro ⟨d, hd, ha⟩ exact ha.trans_lt (add_lt_add_left hd b) +theorem add_le_iff {a b c : Ordinal} (hb : b ≠ 0) : a + b ≤ c ↔ ∀ d < b, a + d < c := by + simpa using (lt_add_iff hb).not + +@[deprecated add_le_iff (since := "2024-12-08")] +theorem add_le_of_forall_add_lt {a b c : Ordinal} (hb : 0 < b) (h : ∀ d < b, a + d < c) : + a + b ≤ c := + (add_le_iff hb.ne').2 h + theorem isLimit_sub {a b} (ha : IsLimit a) (h : b < a) : IsLimit (a - b) := by rw [isLimit_iff, Ordinal.sub_ne_zero_iff_lt, isSuccPrelimit_iff_succ_lt] refine ⟨h, fun c hc ↦ ?_⟩ @@ -2456,18 +2465,6 @@ theorem add_mul_succ {a b : Ordinal} (c) (ba : b + a = a) : (a + b) * succ c = a theorem add_mul_limit {a b c : Ordinal} (ba : b + a = a) (l : IsLimit c) : (a + b) * c = a * c := add_mul_limit_aux ba l fun c' _ => add_mul_succ c' ba -theorem add_le_of_forall_add_lt {a b c : Ordinal} (hb : 0 < b) (h : ∀ d < b, a + d < c) : - a + b ≤ c := by - have H : a + (c - a) = c := - Ordinal.add_sub_cancel_of_le - (by - rw [← add_zero a] - exact (h _ hb).le) - rw [← H] - apply add_le_add_left _ a - by_contra! hb - exact (h _ hb).ne H - theorem IsNormal.apply_omega0 {f : Ordinal.{u} → Ordinal.{v}} (hf : IsNormal f) : ⨆ n : ℕ, f n = f ω := by rw [← iSup_natCast, hf.map_iSup] From 28a2a67063eb1ad7786c5237a793c130faa15bec Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 10 Dec 2024 16:18:09 +0000 Subject: [PATCH 640/829] feat: isomorphism between the non-divisors and units of a group with zero (#19868) From FLT --- .../GroupWithZero/NonZeroDivisors.lean | 29 +++++++++++++------ 1 file changed, 20 insertions(+), 9 deletions(-) diff --git a/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean b/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean index e1c973a0dcb5f..1cae6fca2dc4b 100644 --- a/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean +++ b/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean @@ -105,10 +105,9 @@ def nonZeroSMulDivisors (R : Type*) [MonoidWithZero R] (M : Type _) [Zero M] [Mu /-- The notation for the submonoid of non-zero smul-divisors. -/ scoped[nonZeroSMulDivisors] notation:9000 R "⁰[" M "]" => nonZeroSMulDivisors R M -section nonZeroDivisors - open nonZeroDivisors +section MonoidWithZero variable {M M' M₁ R R' F : Type*} [MonoidWithZero M] [MonoidWithZero M'] [CommMonoidWithZero M₁] [Ring R] [CommRing R'] @@ -180,11 +179,6 @@ theorem mul_mem_nonZeroDivisors {a b : M₁} : a * b ∈ M₁⁰ ↔ a ∈ M₁ apply hb rw [mul_assoc, hx] -theorem isUnit_of_mem_nonZeroDivisors {G₀ : Type*} [GroupWithZero G₀] {x : G₀} - (hx : x ∈ nonZeroDivisors G₀) : IsUnit x := - ⟨⟨x, x⁻¹, mul_inv_cancel₀ (nonZeroDivisors.ne_zero hx), - inv_mul_cancel₀ (nonZeroDivisors.ne_zero hx)⟩, rfl⟩ - lemma IsUnit.mem_nonZeroDivisors {a : M} (ha : IsUnit a) : a ∈ M⁰ := fun _ h ↦ ha.mul_left_eq_zero.mp h @@ -199,7 +193,7 @@ theorem eq_zero_of_ne_zero_of_mul_left_eq_zero [NoZeroDivisors M] {x y : M} (hnx theorem mem_nonZeroDivisors_of_ne_zero [NoZeroDivisors M] {x : M} (hx : x ≠ 0) : x ∈ M⁰ := fun _ ↦ eq_zero_of_ne_zero_of_mul_right_eq_zero hx -theorem mem_nonZeroDivisors_iff_ne_zero [NoZeroDivisors M] [Nontrivial M] {x : M} : +@[simp] lemma mem_nonZeroDivisors_iff_ne_zero [NoZeroDivisors M] [Nontrivial M] {x : M} : x ∈ M⁰ ↔ x ≠ 0 := ⟨nonZeroDivisors.ne_zero, mem_nonZeroDivisors_of_ne_zero⟩ variable [FunLike F M M'] @@ -252,7 +246,24 @@ lemma isUnit_iff_mem_nonZeroDivisors_of_finite [Finite R] {a : R} : rw [← sub_eq_zero, ← sub_mul] at hbc exact sub_eq_zero.mp (ha _ hbc) -end nonZeroDivisors +end MonoidWithZero + +section GroupWithZero +variable {G₀ : Type*} [GroupWithZero G₀] {x : G₀} + +/-- Canonical isomorphism between the non-zero-divisors and units of a group with zero. -/ +@[simps] +noncomputable def nonZeroDivisorsEquivUnits : G₀⁰ ≃* G₀ˣ where + toFun u := .mk0 _ <| mem_nonZeroDivisors_iff_ne_zero.1 u.2 + invFun u := ⟨u, u.isUnit.mem_nonZeroDivisors⟩ + left_inv u := rfl + right_inv u := by simp + map_mul' u v := by simp + +lemma isUnit_of_mem_nonZeroDivisors (hx : x ∈ nonZeroDivisors G₀) : IsUnit x := + (nonZeroDivisorsEquivUnits ⟨x, hx⟩).isUnit + +end GroupWithZero section nonZeroSMulDivisors From 02c577686ffb56d20d204b4cc09d86f76d2aff32 Mon Sep 17 00:00:00 2001 From: Jack Valmadre Date: Tue, 10 Dec 2024 16:45:08 +0000 Subject: [PATCH 641/829] feat(Analysis/Distribution): generalize linearity of SchwartzMap.bilinLeftCLM (#19823) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Generalize SchwartzMap.bilinLeftCLM from ℝ-linearity to 𝕜-linearity with `[NontriviallyNormedField 𝕜]`. Add ContinuousLinearMap.bilinRestrictScalars for convenience. --- .../Analysis/Distribution/SchwartzSpace.lean | 20 ++++++----- .../NormedSpace/OperatorNorm/Bilinear.lean | 33 +++++++++++++++++++ 2 files changed, 45 insertions(+), 8 deletions(-) diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index 40e9573a8eb34..c5823ee573914 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -755,14 +755,15 @@ end EvalCLM section Multiplication -variable [NormedAddCommGroup D] [NormedSpace ℝ D] -variable [NormedAddCommGroup G] [NormedSpace ℝ G] +variable [NontriviallyNormedField 𝕜] [NormedAlgebra ℝ 𝕜] + [NormedAddCommGroup D] [NormedSpace ℝ D] + [NormedAddCommGroup G] [NormedSpace ℝ G] + [NormedSpace 𝕜 E] [NormedSpace 𝕜 F] [NormedSpace 𝕜 G] /-- The map `f ↦ (x ↦ B (f x) (g x))` as a continuous `𝕜`-linear map on Schwartz space, where `B` is a continuous `𝕜`-linear map and `g` is a function of temperate growth. -/ -def bilinLeftCLM (B : E →L[ℝ] F →L[ℝ] G) {g : D → F} (hg : g.HasTemperateGrowth) : - 𝓢(D, E) →L[ℝ] 𝓢(D, G) := by - -- Todo (after port): generalize to `B : E →L[𝕜] F →L[𝕜] G` and `𝕜`-linear +def bilinLeftCLM (B : E →L[𝕜] F →L[𝕜] G) {g : D → F} (hg : g.HasTemperateGrowth) : + 𝓢(D, E) →L[𝕜] 𝓢(D, G) := by refine mkCLM (fun f x => B (f x) (g x)) (fun _ _ _ => by simp only [map_add, add_left_inj, Pi.add_apply, eq_self_iff_true, @@ -770,7 +771,8 @@ def bilinLeftCLM (B : E →L[ℝ] F →L[ℝ] G) {g : D → F} (hg : g.HasTemper (fun _ _ _ => by simp only [smul_apply, map_smul, ContinuousLinearMap.coe_smul', Pi.smul_apply, RingHom.id_apply]) - (fun f => (B.isBoundedBilinearMap.contDiff.restrict_scalars ℝ).comp (f.smooth'.prod hg.1)) ?_ + (fun f => (B.bilinearRestrictScalars ℝ).isBoundedBilinearMap.contDiff.comp + (f.smooth'.prod hg.1)) ?_ rintro ⟨k, n⟩ rcases hg.norm_iteratedFDeriv_le_uniform_aux n with ⟨l, C, hC, hgrowth⟩ use @@ -778,10 +780,12 @@ def bilinLeftCLM (B : E →L[ℝ] F →L[ℝ] G) {g : D → F} (hg : g.HasTemper by positivity intro f x have hxk : 0 ≤ ‖x‖ ^ k := by positivity + simp_rw [← ContinuousLinearMap.bilinearRestrictScalars_apply_apply ℝ B] have hnorm_mul := - ContinuousLinearMap.norm_iteratedFDeriv_le_of_bilinear B f.smooth' hg.1 x (n := n) - (mod_cast le_top) + ContinuousLinearMap.norm_iteratedFDeriv_le_of_bilinear (B.bilinearRestrictScalars ℝ) + f.smooth' hg.1 x (n := n) (mod_cast le_top) refine le_trans (mul_le_mul_of_nonneg_left hnorm_mul hxk) ?_ + rw [ContinuousLinearMap.norm_bilinearRestrictScalars] move_mul [← ‖B‖] simp_rw [mul_assoc ‖B‖] gcongr _ * ?_ diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean index 61f5199b50531..6b959b42931e5 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean @@ -460,3 +460,36 @@ theorem norm_smulRightL_apply (c : E →L[𝕜] 𝕜) (f : Fₗ) : ‖smulRightL end ContinuousLinearMap end SemiNormed + +section Restrict + +namespace ContinuousLinearMap + +variable {𝕜' : Type*} [NontriviallyNormedField 𝕜] [NontriviallyNormedField 𝕜'] [NormedAlgebra 𝕜 𝕜'] + [SeminormedAddCommGroup E] [NormedSpace 𝕜 E] [NormedSpace 𝕜' E] [IsScalarTower 𝕜 𝕜' E] + [SeminormedAddCommGroup F] [NormedSpace 𝕜 F] [NormedSpace 𝕜' F] [IsScalarTower 𝕜 𝕜' F] + [SeminormedAddCommGroup G] [NormedSpace 𝕜 G] [NormedSpace 𝕜' G] [IsScalarTower 𝕜 𝕜' G] + +variable (𝕜) in +/-- Convenience function for restricting the linearity of a bilinear map. -/ +def bilinearRestrictScalars (B : E →L[𝕜'] F →L[𝕜'] G) : E →L[𝕜] F →L[𝕜] G := + (restrictScalarsL 𝕜' F G 𝕜 𝕜).comp (B.restrictScalars 𝕜) + +variable (B : E →L[𝕜'] F →L[𝕜'] G) (x : E) (y : F) + +theorem bilinearRestrictScalars_eq_restrictScalarsL_comp_restrictScalars : + B.bilinearRestrictScalars 𝕜 = (restrictScalarsL 𝕜' F G 𝕜 𝕜).comp (B.restrictScalars 𝕜) := rfl + +theorem bilinearRestrictScalars_eq_restrictScalars_restrictScalarsL_comp : + B.bilinearRestrictScalars 𝕜 = restrictScalars 𝕜 ((restrictScalarsL 𝕜' F G 𝕜 𝕜').comp B) := rfl + +variable (𝕜) in +@[simp] +theorem bilinearRestrictScalars_apply_apply : (B.bilinearRestrictScalars 𝕜) x y = B x y := rfl + +@[simp] +theorem norm_bilinearRestrictScalars : ‖B.bilinearRestrictScalars 𝕜‖ = ‖B‖ := rfl + +end ContinuousLinearMap + +end Restrict From a38f191d0b038365a318f3334a591b52423e5aad Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 10 Dec 2024 17:02:58 +0000 Subject: [PATCH 642/829] feat: sup-closed sets are closed under finite suprema (#18990) From GrowthInGroups --- Mathlib/Order/SupClosed.lean | 66 +++++++++++++++++++++++++++++++++++- 1 file changed, 65 insertions(+), 1 deletion(-) diff --git a/Mathlib/Order/SupClosed.lean b/Mathlib/Order/SupClosed.lean index 84fdcccc76f9d..fdec9d591598c 100644 --- a/Mathlib/Order/SupClosed.lean +++ b/Mathlib/Order/SupClosed.lean @@ -7,6 +7,7 @@ import Mathlib.Data.Finset.Lattice.Fold import Mathlib.Data.Finset.Powerset import Mathlib.Data.Set.Finite.Basic import Mathlib.Order.Closure +import Mathlib.Order.ConditionallyCompleteLattice.Finset /-! # Sets closed under join/meet @@ -29,7 +30,7 @@ is automatically complete. All dually for `⊓`. greatest lower bound is automatically complete. -/ -variable {F α β : Type*} +variable {ι : Sort*} {F α β : Type*} section SemilatticeSup variable [SemilatticeSup α] [SemilatticeSup β] @@ -501,3 +502,66 @@ def SemilatticeInf.toCompleteSemilatticeInf [SemilatticeInf α] (sInf : Set α sInf := fun s => sInf (infClosure s) sInf_le _ _ ha := (h _ infClosed_infClosure).1 <| subset_infClosure ha le_sInf s a ha := (le_isGLB_iff <| h _ infClosed_infClosure).2 <| by rwa [lowerBounds_infClosure] + + +section ConditionallyCompleteLattice +variable [ConditionallyCompleteLattice α] {f : ι → α} {s t : Set α} + +lemma SupClosed.iSup_mem_of_nonempty [Finite ι] [Nonempty ι] (hs : SupClosed s) + (hf : ∀ i, f i ∈ s) : ⨆ i, f i ∈ s := by + cases nonempty_fintype (PLift ι) + rw [← iSup_plift_down, ← Finset.sup'_univ_eq_ciSup] + exact hs.finsetSup'_mem Finset.univ_nonempty fun _ _ ↦ hf _ + +lemma InfClosed.iInf_mem_of_nonempty [Finite ι] [Nonempty ι] (hs : InfClosed s) + (hf : ∀ i, f i ∈ s) : ⨅ i, f i ∈ s := hs.dual.iSup_mem_of_nonempty hf + +lemma SupClosed.sSup_mem_of_nonempty (hs : SupClosed s) (ht : t.Finite) (ht' : t.Nonempty) + (hts : t ⊆ s) : sSup t ∈ s := by + have := ht.to_subtype + have := ht'.to_subtype + rw [sSup_eq_iSup'] + exact hs.iSup_mem_of_nonempty (by simpa) + +lemma InfClosed.sInf_mem_of_nonempty (hs : InfClosed s) (ht : t.Finite) (ht' : t.Nonempty) + (hts : t ⊆ s) : sInf t ∈ s := hs.dual.sSup_mem_of_nonempty ht ht' hts + +end ConditionallyCompleteLattice + +variable [CompleteLattice α] {f : ι → α} {s t : Set α} + +lemma SupClosed.biSup_mem_of_nonempty {ι : Type*} {t : Set ι} {f : ι → α} (hs : SupClosed s) + (ht : t.Finite) (ht' : t.Nonempty) (hf : ∀ i ∈ t, f i ∈ s) : ⨆ i ∈ t, f i ∈ s := by + rw [← sSup_image] + exact hs.sSup_mem_of_nonempty (ht.image _) (by simpa) (by simpa) + +lemma InfClosed.biInf_mem_of_nonempty {ι : Type*} {t : Set ι} {f : ι → α} (hs : InfClosed s) + (ht : t.Finite) (ht' : t.Nonempty) (hf : ∀ i ∈ t, f i ∈ s) : ⨅ i ∈ t, f i ∈ s := + hs.dual.biSup_mem_of_nonempty ht ht' hf + +lemma SupClosed.iSup_mem [Finite ι] (hs : SupClosed s) (hbot : ⊥ ∈ s) (hf : ∀ i, f i ∈ s) : + ⨆ i, f i ∈ s := by + cases isEmpty_or_nonempty ι + · simpa [iSup_of_empty] + · exact hs.iSup_mem_of_nonempty hf + +lemma InfClosed.iInf_mem [Finite ι] (hs : InfClosed s) (htop : ⊤ ∈ s) (hf : ∀ i, f i ∈ s) : + ⨅ i, f i ∈ s := hs.dual.iSup_mem htop hf + +lemma SupClosed.sSup_mem (hs : SupClosed s) (ht : t.Finite) (hbot : ⊥ ∈ s) (hts : t ⊆ s) : + sSup t ∈ s := by + have := ht.to_subtype + rw [sSup_eq_iSup'] + exact hs.iSup_mem hbot (by simpa) + +lemma InfClosed.sInf_mem (hs : InfClosed s) (ht : t.Finite) (htop : ⊤ ∈ s) (hts : t ⊆ s) : + sInf t ∈ s := hs.dual.sSup_mem ht htop hts + +lemma SupClosed.biSup_mem {ι : Type*} {t : Set ι} {f : ι → α} (hs : SupClosed s) + (ht : t.Finite) (hbot : ⊥ ∈ s) (hf : ∀ i ∈ t, f i ∈ s) : ⨆ i ∈ t, f i ∈ s := by + rw [← sSup_image] + exact hs.sSup_mem (ht.image _) hbot (by simpa) + +lemma InfClosed.biInf_mem {ι : Type*} {t : Set ι} {f : ι → α} (hs : InfClosed s) + (ht : t.Finite) (htop : ⊤ ∈ s) (hf : ∀ i ∈ t, f i ∈ s) : ⨅ i ∈ t, f i ∈ s := + hs.dual.biSup_mem ht htop hf From c3b16e029e0572ec8bfd4f21ce79a391eb59ec50 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 10 Dec 2024 17:03:00 +0000 Subject: [PATCH 643/829] chore: make `Nat.cast_smul_eq_nsmul`/`Int.cast_smul_eq_zsmul` `norm_cast` lemmas (#19869) --- Mathlib/Algebra/Module/End.lean | 1 + Mathlib/Algebra/Module/NatInt.lean | 1 + 2 files changed, 2 insertions(+) diff --git a/Mathlib/Algebra/Module/End.lean b/Mathlib/Algebra/Module/End.lean index 405cb931c87f9..d11fb9a0448a3 100644 --- a/Mathlib/Algebra/Module/End.lean +++ b/Mathlib/Algebra/Module/End.lean @@ -81,6 +81,7 @@ section variable (R) /-- `zsmul` is equal to any other module structure via a cast. -/ +@[norm_cast] lemma Int.cast_smul_eq_zsmul (n : ℤ) (b : M) : (n : R) • b = n • b := have : ((smulAddHom R M).flip b).comp (Int.castAddHom R) = (smulAddHom ℤ M).flip b := by apply AddMonoidHom.ext_int diff --git a/Mathlib/Algebra/Module/NatInt.lean b/Mathlib/Algebra/Module/NatInt.lean index d1aab097a652d..22860f1391265 100644 --- a/Mathlib/Algebra/Module/NatInt.lean +++ b/Mathlib/Algebra/Module/NatInt.lean @@ -95,6 +95,7 @@ section variable (R) /-- `nsmul` is equal to any other module structure via a cast. -/ +@[norm_cast] lemma Nat.cast_smul_eq_nsmul (n : ℕ) (b : M) : (n : R) • b = n • b := by induction n with | zero => rw [Nat.cast_zero, zero_smul, zero_smul] From 57bf40112b5f3e0f11f619d4d67fe2de66f9e650 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 10 Dec 2024 17:35:30 +0000 Subject: [PATCH 644/829] =?UTF-8?q?feat:=20lower=20bound=20on=20the=20Ruzs?= =?UTF-8?q?a-Szemer=C3=A9di=20problem=20(#19000)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From LeanCamCombi --- Mathlib.lean | 1 + .../Extremal/RuzsaSzemeredi.lean | 269 ++++++++++++++++++ 2 files changed, 270 insertions(+) create mode 100644 Mathlib/Combinatorics/Extremal/RuzsaSzemeredi.lean diff --git a/Mathlib.lean b/Mathlib.lean index b5020b3a7fbb0..c4527bf18ce4c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2199,6 +2199,7 @@ import Mathlib.Combinatorics.Enumerative.DyckWord import Mathlib.Combinatorics.Enumerative.IncidenceAlgebra import Mathlib.Combinatorics.Enumerative.InclusionExclusion import Mathlib.Combinatorics.Enumerative.Partition +import Mathlib.Combinatorics.Extremal.RuzsaSzemeredi import Mathlib.Combinatorics.HalesJewett import Mathlib.Combinatorics.Hall.Basic import Mathlib.Combinatorics.Hall.Finite diff --git a/Mathlib/Combinatorics/Extremal/RuzsaSzemeredi.lean b/Mathlib/Combinatorics/Extremal/RuzsaSzemeredi.lean new file mode 100644 index 0000000000000..59c4a36c6596f --- /dev/null +++ b/Mathlib/Combinatorics/Extremal/RuzsaSzemeredi.lean @@ -0,0 +1,269 @@ +/- +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.Combinatorics.Additive.AP.Three.Behrend +import Mathlib.Combinatorics.SimpleGraph.Triangle.Tripartite +import Mathlib.Tactic.Rify + +/-! +# The Ruzsa-Szemerédi problem + +This file proves the lower bound of the Ruzsa-Szemerédi problem. The problem is to find the maximum +number of edges that a graph on `n` vertices can have if all edges belong to at most one triangle. + +The lower bound comes from turning the big 3AP-free set from Behrend's construction into a graph +that has the property that every triangle gives a (possibly trivial) arithmetic progression on the +original set. + +## Main declarations + +* `ruzsaSzemerediNumberNat n`: Maximum number of edges a graph on `n` vertices can have such that + each edge belongs to exactly one triangle. +* `ruzsaSzemerediNumberNat_asymptotic_lower_bound`: There exists a graph with `n` vertices and + `Ω((n ^ 2 * exp (-4 * sqrt (log n))))` edges such that each edge belongs to exactly one triangle. +-/ + +open Finset Nat Real SimpleGraph Sum3 SimpleGraph.TripartiteFromTriangles +open Fintype (card) +open scoped Pointwise + +variable {α β : Type*} + +/-! ### The Ruzsa-Szemerédi number -/ + +section ruzsaSzemerediNumber +variable [DecidableEq α] [DecidableEq β] [Fintype α] [Fintype β] {G H : SimpleGraph α} + +variable (α) in +/-- The **Ruzsa-Szemerédi number** of a fintype is the maximum number of edges a locally linear +graph on that type can have. + +In other words, `ruzsaSzemerediNumber α` is the maximum number of edges a graph on `α` can have such +that each edge belongs to exactly one triangle. -/ +noncomputable def ruzsaSzemerediNumber : ℕ := by + classical + exact Nat.findGreatest (fun m ↦ ∃ (G : SimpleGraph α) (_ : DecidableRel G.Adj), + #(G.cliqueFinset 3) = m ∧ G.LocallyLinear) ((card α).choose 3) + +open scoped Classical in +lemma ruzsaSzemerediNumber_le : ruzsaSzemerediNumber α ≤ (card α).choose 3 := Nat.findGreatest_le _ + +lemma ruzsaSzemerediNumber_spec : + ∃ (G : SimpleGraph α) (_ : DecidableRel G.Adj), + #(G.cliqueFinset 3) = ruzsaSzemerediNumber α ∧ G.LocallyLinear := by + classical + exact @Nat.findGreatest_spec _ + (fun m ↦ ∃ (G : SimpleGraph α) (_ : DecidableRel G.Adj), + #(G.cliqueFinset 3) = m ∧ G.LocallyLinear) _ _ (Nat.zero_le _) + ⟨⊥, inferInstance, by simp, locallyLinear_bot⟩ + +variable {n : ℕ} + +lemma SimpleGraph.LocallyLinear.le_ruzsaSzemerediNumber [DecidableRel G.Adj] + (hG : G.LocallyLinear) : #(G.cliqueFinset 3) ≤ ruzsaSzemerediNumber α := by + classical + exact le_findGreatest card_cliqueFinset_le ⟨G, inferInstance, by congr, hG⟩ + +lemma ruzsaSzemerediNumber_mono (f : α ↪ β) : ruzsaSzemerediNumber α ≤ ruzsaSzemerediNumber β := by + classical + refine findGreatest_mono ?_ (choose_mono _ <| Fintype.card_le_of_embedding f) + rintro n ⟨G, _, rfl, hG⟩ + refine ⟨G.map f, inferInstance, ?_, hG.map _⟩ + rw [← card_map ⟨map f, Finset.map_injective _⟩, ← cliqueFinset_map G f] + decide + +lemma ruzsaSzemerediNumber_congr (e : α ≃ β) : ruzsaSzemerediNumber α = ruzsaSzemerediNumber β := + (ruzsaSzemerediNumber_mono (e : α ↪ β)).antisymm <| ruzsaSzemerediNumber_mono e.symm + +/-- The `n`-th **Ruzsa-Szemerédi number** is the maximum number of edges a locally linear graph on +`n` vertices can have. + +In other words, `ruzsaSzemerediNumberNat n` is the maximum number of edges a graph on `n` vertices +can have such that each edge belongs to exactly one triangle. -/ +noncomputable def ruzsaSzemerediNumberNat (n : ℕ) : ℕ := ruzsaSzemerediNumber (Fin n) + +@[simp] +lemma ruzsaSzemerediNumberNat_card : ruzsaSzemerediNumberNat (card α) = ruzsaSzemerediNumber α := + ruzsaSzemerediNumber_congr (Fintype.equivFin _).symm + +lemma ruzsaSzemerediNumberNat_mono : Monotone ruzsaSzemerediNumberNat := fun _m _n h => + ruzsaSzemerediNumber_mono (Fin.castLEEmb h) + +lemma ruzsaSzemerediNumberNat_le : ruzsaSzemerediNumberNat n ≤ n.choose 3 := + ruzsaSzemerediNumber_le.trans_eq <| by rw [Fintype.card_fin] + +@[simp] lemma ruzsaSzemerediNumberNat_zero : ruzsaSzemerediNumberNat 0 = 0 := + le_zero_iff.1 ruzsaSzemerediNumberNat_le + +@[simp] lemma ruzsaSzemerediNumberNat_one : ruzsaSzemerediNumberNat 1 = 0 := + le_zero_iff.1 ruzsaSzemerediNumberNat_le + +@[simp] lemma ruzsaSzemerediNumberNat_two : ruzsaSzemerediNumberNat 2 = 0 := + le_zero_iff.1 ruzsaSzemerediNumberNat_le + +end ruzsaSzemerediNumber + +/-! ### The Ruzsa-Szemerédi construction -/ + +section RuzsaSzemeredi +variable [Fintype α] [CommRing α] {s : Finset α} {x : α × α × α} + +/-- The triangle indices for the Ruzsa-Szemerédi construction. -/ +private def triangleIndices (s : Finset α) : Finset (α × α × α) := + (univ ×ˢ s).map + ⟨fun xa ↦ (xa.1, xa.1 + xa.2, xa.1 + 2 * xa.2), by + rintro ⟨x, a⟩ ⟨y, b⟩ h + simp only [Prod.ext_iff] at h + obtain rfl := h.1 + obtain rfl := add_right_injective _ h.2.1 + rfl⟩ + +@[simp] +private lemma mem_triangleIndices : + x ∈ triangleIndices s ↔ ∃ y, ∃ a ∈ s, (y, y + a, y + 2 * a) = x := by simp [triangleIndices] + +@[simp] +private lemma card_triangleIndices : #(triangleIndices s) = card α * #s := by + simp [triangleIndices, card_univ] + +private lemma noAccidental (hs : ThreeAPFree (s : Set α)) : + NoAccidental (triangleIndices s : Finset (α × α × α)) where + eq_or_eq_or_eq := by + simp only [mem_triangleIndices, Prod.mk.inj_iff, exists_prop, forall_exists_index, and_imp] + rintro _ _ _ _ _ _ d a ha rfl rfl rfl b' b hb rfl rfl h₁ d' c hc rfl h₂ rfl + have : a + c = b + b := by linear_combination h₁.symm - h₂.symm + obtain rfl := hs ha hb hc this + simp_all + +variable [Fact <| IsUnit (2 : α)] + +private instance : ExplicitDisjoint (triangleIndices s : Finset (α × α × α)) where + inj₀ := by + simp only [mem_triangleIndices, Prod.mk.inj_iff, exists_prop, forall_exists_index, and_imp] + rintro _ _ _ _ x a ha rfl rfl rfl y b hb rfl h₁ h₂ + linear_combination 2 * h₁.symm - h₂.symm + inj₁ := by + simp only [mem_triangleIndices, Prod.mk.inj_iff, exists_prop, forall_exists_index, and_imp] + rintro _ _ _ _ x a ha rfl rfl rfl y b hb rfl rfl h + simpa [(Fact.out (p := IsUnit (2 : α))).mul_right_inj, eq_comm] using h + inj₂ := by + simp only [mem_triangleIndices, Prod.mk.inj_iff, exists_prop, forall_exists_index, and_imp] + rintro _ _ _ _ x a ha rfl rfl rfl y b hb rfl h rfl + simpa [(Fact.out (p := IsUnit (2 : α))).mul_right_inj, eq_comm] using h + +private lemma locallyLinear (hs : ThreeAPFree (s : Set α)) : + (graph <| triangleIndices s).LocallyLinear := + haveI := noAccidental hs; TripartiteFromTriangles.locallyLinear _ + +private lemma card_edgeFinset (hs : ThreeAPFree (s : Set α)) [DecidableEq α] : + #(graph <| triangleIndices s).edgeFinset = 3 * card α * #s := by + haveI := noAccidental hs + rw [(locallyLinear hs).card_edgeFinset, card_triangles, card_triangleIndices, mul_assoc] + +end RuzsaSzemeredi + +variable (α) [Fintype α] [DecidableEq α] [CommRing α] [Fact <| IsUnit (2 : α)] + +lemma addRothNumber_le_ruzsaSzemerediNumber : + card α * addRothNumber (univ : Finset α) ≤ ruzsaSzemerediNumber (Sum α (Sum α α)) := by + obtain ⟨s, -, hscard, hs⟩ := addRothNumber_spec (univ : Finset α) + haveI := noAccidental hs + rw [← hscard, ← card_triangleIndices, ← card_triangles] + exact (locallyLinear hs).le_ruzsaSzemerediNumber + +lemma rothNumberNat_le_ruzsaSzemerediNumberNat (n : ℕ) : + (2 * n + 1) * rothNumberNat n ≤ ruzsaSzemerediNumberNat (6 * n + 3) := by + let α := Fin (2 * n + 1) + have : Nat.Coprime 2 (2 * n + 1) := by simp + haveI : Fact (IsUnit (2 : Fin (2 * n + 1))) := ⟨by simpa using (ZMod.unitOfCoprime 2 this).isUnit⟩ + calc + (2 * n + 1) * rothNumberNat n + _ = Fintype.card α * addRothNumber (Iio (n : α)) := by + rw [Fin.addRothNumber_eq_rothNumberNat le_rfl, Fintype.card_fin] + _ ≤ Fintype.card α * addRothNumber (univ : Finset α) := by + gcongr; exact subset_univ _ + _ ≤ ruzsaSzemerediNumber (Sum α (Sum α α)) := addRothNumber_le_ruzsaSzemerediNumber _ + _ = ruzsaSzemerediNumberNat (6 * n + 3) := by + simp_rw [← ruzsaSzemerediNumberNat_card, Fintype.card_sum, α, Fintype.card_fin] + ring_nf + +/-- Lower bound on the **Ruzsa-Szemerédi problem** in terms of 3AP-free sets. + +If there exists a 3AP-free subset of `[1, ..., (n - 3) / 6]` of size `m`, then there exists a graph +with `n` vertices and `(n / 3 - 2) * m` edges such that each edge belongs to exactly one triangle. +-/ +theorem rothNumberNat_le_ruzsaSzemerediNumberNat' : + ∀ n : ℕ, (n / 3 - 2 : ℝ) * rothNumberNat ((n - 3) / 6) ≤ ruzsaSzemerediNumberNat n + | 0 => by simp + | 1 => by simp + | 2 => by simp + | n + 3 => by + calc + _ ≤ (↑(2 * (n / 6) + 1) : ℝ) * rothNumberNat (n / 6) := + mul_le_mul_of_nonneg_right ?_ (Nat.cast_nonneg _) + _ ≤ (ruzsaSzemerediNumberNat (6 * (n / 6) + 3) : ℝ) := ?_ + _ ≤ _ := + Nat.cast_le.2 (ruzsaSzemerediNumberNat_mono <| add_le_add_right (Nat.mul_div_le _ _) _) + · norm_num + rw [← div_add_one (three_ne_zero' ℝ), ← le_sub_iff_add_le, div_le_iff₀ (zero_lt_three' ℝ), + add_assoc, add_sub_assoc, add_mul, mul_right_comm] + norm_num + norm_cast + rw [← mul_add_one] + exact (Nat.lt_mul_div_succ _ <| by norm_num).le + · norm_cast + exact rothNumberNat_le_ruzsaSzemerediNumberNat _ + +/-- Explicit lower bound on the **Ruzsa-Szemerédi problem**. + +There exists a graph with `n` vertices and +`(n / 3 - 2) * (n - 3) / 6 * exp (-4 * sqrt (log ((n - 3) / 6)))` edges such that each edge belongs +to exactly one triangle. -/ +theorem ruzsaSzemerediNumberNat_lower_bound (n : ℕ) : + (n / 3 - 2 : ℝ) * ↑((n - 3) / 6) * exp (-4 * sqrt (log ↑((n - 3) / 6))) ≤ + ruzsaSzemerediNumberNat n := by + rw [mul_assoc] + obtain hn | hn := le_total (n / 3 - 2 : ℝ) 0 + · exact (mul_nonpos_of_nonpos_of_nonneg hn <| by positivity).trans (Nat.cast_nonneg _) + exact + (mul_le_mul_of_nonneg_left Behrend.roth_lower_bound hn).trans + (rothNumberNat_le_ruzsaSzemerediNumberNat' _) + +open Asymptotics Filter + +/-- Asymptotic lower bound on the **Ruzsa-Szemerédi problem**. + +There exists a graph with `n` vertices and `Ω((n ^ 2 * exp (-4 * sqrt (log n))))` edges such that +each edge belongs to exactly one triangle. -/ +theorem ruzsaSzemerediNumberNat_asymptotic_lower_bound : + (fun n ↦ n ^ 2 * exp (-4 * sqrt (log n)) : ℕ → ℝ) =O[atTop] + fun n ↦ (ruzsaSzemerediNumberNat n : ℝ) := by + trans fun n ↦ (n / 3 - 2) * ↑((n - 3) / 6) * exp (-4 * sqrt (log ↑((n - 3) / 6))) + · simp_rw [sq] + refine (IsBigO.mul ?_ ?_).mul ?_ + · trans fun n ↦ n / 3 + · simp_rw [div_eq_inv_mul] + exact (isBigO_refl ..).const_mul_right (by norm_num) + refine IsLittleO.right_isBigO_sub ?_ + simpa [div_eq_inv_mul, Function.comp_def] using + .atTop_of_const_mul zero_lt_three (by simp [tendsto_natCast_atTop_atTop]) + · rw [IsBigO_def] + refine ⟨12, ?_⟩ + simp only [IsBigOWith, norm_natCast, eventually_atTop] + exact ⟨15, fun x hx ↦ by norm_cast; omega⟩ + · rw [isBigO_exp_comp_exp_comp] + refine ⟨0, ?_⟩ + simp only [neg_mul, eventually_map, Pi.sub_apply, sub_neg_eq_add, neg_add_le_iff_le_add, + add_zero, ofNat_pos, _root_.mul_le_mul_left, eventually_atTop] + refine ⟨9, fun x hx ↦ ?_⟩ + gcongr + · simp + omega + · omega + · refine .of_bound 1 ?_ + simp only [neg_mul, norm_eq_abs, norm_natCast, one_mul, eventually_atTop] + refine ⟨6, fun n hn ↦ ?_⟩ + have : (0 : ℝ) ≤ n / 3 - 2 := by rify at hn; linarith + simpa using abs_le_abs_of_nonneg (by positivity) (ruzsaSzemerediNumberNat_lower_bound n) From 7817ae82d4ca6c4d0399b897271dc8a26c287cf8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 10 Dec 2024 17:35:31 +0000 Subject: [PATCH 645/829] feat: relation of covering by cosets (#19509) Define a predicate for a set to be covered by at most `K` cosets of another set. This is a fundamental relation to study in additive combinatorics. From GrowthInGroups (LeanCamCombi) --- Mathlib.lean | 1 + Mathlib/Combinatorics/Additive/CovBySMul.lean | 70 +++++++++++++++++++ 2 files changed, 71 insertions(+) create mode 100644 Mathlib/Combinatorics/Additive/CovBySMul.lean diff --git a/Mathlib.lean b/Mathlib.lean index c4527bf18ce4c..8abf0cd53e010 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2174,6 +2174,7 @@ import Mathlib.Combinatorics.Additive.AP.Three.Defs import Mathlib.Combinatorics.Additive.CauchyDavenport import Mathlib.Combinatorics.Additive.Corner.Defs import Mathlib.Combinatorics.Additive.Corner.Roth +import Mathlib.Combinatorics.Additive.CovBySMul import Mathlib.Combinatorics.Additive.Dissociation import Mathlib.Combinatorics.Additive.DoublingConst import Mathlib.Combinatorics.Additive.ETransform diff --git a/Mathlib/Combinatorics/Additive/CovBySMul.lean b/Mathlib/Combinatorics/Additive/CovBySMul.lean new file mode 100644 index 0000000000000..d85c2f8f8d894 --- /dev/null +++ b/Mathlib/Combinatorics/Additive/CovBySMul.lean @@ -0,0 +1,70 @@ +/- +Copyright (c) 2024 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import Mathlib.Algebra.Group.Pointwise.Finset.Basic +import Mathlib.Data.Real.Basic + +/-! +# Relation of covering by cosets + +This file defines a predicate for a set to be covered by at most `K` cosets of another set. + +This is a fundamental relation to study in additive combinatorics. +-/ + +open scoped Finset Pointwise + +variable {M N X : Type*} [Monoid M] [Monoid N] [MulAction M X] [MulAction N X] {K L : ℝ} + {A A₁ A₂ B B₁ B₂ C : Set X} + +variable (M) in +/-- Predicate for a set `A` to be covered by at most `K` cosets of another set `B` under the action +by the monoid `M`. -/ +@[to_additive "Predicate for a set `A` to be covered by at most `K` cosets of another set `B` under +the action by the monoid `M`."] +def CovBySMul (K : ℝ) (A B : Set X) : Prop := ∃ F : Finset M, #F ≤ K ∧ A ⊆ (F : Set M) • B + +@[to_additive (attr := simp, refl)] +lemma CovBySMul.rfl : CovBySMul M 1 A A := ⟨1, by simp⟩ + +@[to_additive (attr := simp)] +lemma CovBySMul.of_subset (hAB : A ⊆ B) : CovBySMul M 1 A B := ⟨1, by simpa⟩ + +@[to_additive] lemma CovBySMul.nonneg : CovBySMul M K A B → 0 ≤ K := by + rintro ⟨F, hF, -⟩; exact (#F).cast_nonneg.trans hF + +@[to_additive (attr := simp)] +lemma covBySMul_zero : CovBySMul M 0 A B ↔ A = ∅ := by simp [CovBySMul] + +@[to_additive] +lemma CovBySMul.mono (hKL : K ≤ L) : CovBySMul M K A B → CovBySMul M L A B := by + rintro ⟨F, hF, hFAB⟩; exact ⟨F, hF.trans hKL, hFAB⟩ + +@[to_additive] lemma CovBySMul.trans [MulAction M N] [IsScalarTower M N X] + (hAB : CovBySMul M K A B) (hBC : CovBySMul N L B C) : CovBySMul N (K * L) A C := by + classical + have := hAB.nonneg + obtain ⟨F₁, hF₁, hFAB⟩ := hAB + obtain ⟨F₂, hF₂, hFBC⟩ := hBC + refine ⟨F₁ • F₂, ?_, ?_⟩ + · calc + (#(F₁ • F₂) : ℝ) ≤ #F₁ * #F₂ := mod_cast Finset.card_smul_le + _ ≤ K * L := by gcongr + · calc + A ⊆ (F₁ : Set M) • B := hFAB + _ ⊆ (F₁ : Set M) • (F₂ : Set N) • C := by gcongr + _ = (↑(F₁ • F₂) : Set N) • C := by simp + +@[to_additive] +lemma CovBySMul.subset_left (hA : A₁ ⊆ A₂) (hAB : CovBySMul M K A₂ B) : + CovBySMul M K A₁ B := by simpa using (CovBySMul.of_subset (M := M) hA).trans hAB + +@[to_additive] +lemma CovBySMul.subset_right (hB : B₁ ⊆ B₂) (hAB : CovBySMul M K A B₁) : + CovBySMul M K A B₂ := by simpa using hAB.trans (.of_subset (M := M) hB) + +@[to_additive] +lemma CovBySMul.subset (hA : A₁ ⊆ A₂) (hB : B₁ ⊆ B₂) (hAB : CovBySMul M K A₂ B₁) : + CovBySMul M K A₁ B₂ := (hAB.subset_left hA).subset_right hB From 16956ef06e49e3e4f34d7be6d97f170580b42a79 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 10 Dec 2024 17:35:33 +0000 Subject: [PATCH 646/829] feat(Combinatorics/Additive): characterise sets with no doubling (#19513) They are exactly the empty set and cosets of a subgroup. From LeanCamCombi Co-authored-by: Patrick Luo --- Mathlib.lean | 1 + .../Algebra/Group/Pointwise/Finset/Basic.lean | 6 +- .../Additive/VerySmallDoubling.lean | 62 +++++++++++++++++++ 3 files changed, 67 insertions(+), 2 deletions(-) create mode 100644 Mathlib/Combinatorics/Additive/VerySmallDoubling.lean diff --git a/Mathlib.lean b/Mathlib.lean index 8abf0cd53e010..c5d560bd39014 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2185,6 +2185,7 @@ import Mathlib.Combinatorics.Additive.PluenneckeRuzsa import Mathlib.Combinatorics.Additive.Randomisation import Mathlib.Combinatorics.Additive.RuzsaCovering import Mathlib.Combinatorics.Additive.SmallTripling +import Mathlib.Combinatorics.Additive.VerySmallDoubling import Mathlib.Combinatorics.Colex import Mathlib.Combinatorics.Configuration import Mathlib.Combinatorics.Derangements.Basic diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index 2e78e81394ffc..a01f060a58f0e 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -740,11 +740,11 @@ theorem singleton_div_singleton (a b : α) : ({a} : Finset α) / {b} = {a / b} : theorem div_subset_div : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ / t₁ ⊆ s₂ / t₂ := image₂_subset -@[to_additive] +@[to_additive (attr := gcongr)] theorem div_subset_div_left : t₁ ⊆ t₂ → s / t₁ ⊆ s / t₂ := image₂_subset_left -@[to_additive] +@[to_additive (attr := gcongr)] theorem div_subset_div_right : s₁ ⊆ s₂ → s₁ / t ⊆ s₂ / t := image₂_subset_right @@ -1472,6 +1472,8 @@ section Mul variable [Mul α] [DecidableEq α] {s t u : Finset α} {a : α} +@[to_additive] lemma smul_finset_subset_mul : a ∈ s → a • t ⊆ s * t := image_subset_image₂_right + @[to_additive] theorem op_smul_finset_subset_mul : a ∈ t → op a • s ⊆ s * t := image_subset_image₂_left diff --git a/Mathlib/Combinatorics/Additive/VerySmallDoubling.lean b/Mathlib/Combinatorics/Additive/VerySmallDoubling.lean new file mode 100644 index 0000000000000..b6015f656048d --- /dev/null +++ b/Mathlib/Combinatorics/Additive/VerySmallDoubling.lean @@ -0,0 +1,62 @@ +/- +Copyright (c) 2024 Yaël Dillies, Patrick Luo. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies, Patrick Luo +-/ +import Mathlib.Algebra.Group.Pointwise.Finset.Basic +import Mathlib.GroupTheory.GroupAction.Defs + +/-! +# Sets with very small doubling + +This file characterises sets with no doubling (finsets `A` such that `#(A ^ 2) = #A`) as the sets +which are either empty or translates of a subgroup. + +For the converse, use the existing facts from the pointwise API: `∅ ^ 2 = ∅` (`Finset.empty_pow`), +`(a • H) ^ 2 = a ^ 2 • H ^ 2 = a ^ 2 • H` (`smul_pow`, `coe_set_pow`) + +## TODO + +* Do we need a version stated using the doubling constant (`Finset.mulConst`)? +* Add characterisation for sets with doubling < 3/2 +-/ + +open MulOpposite MulAction +open scoped Pointwise RightActions + +namespace Finset +variable {G : Type*} [Group G] [DecidableEq G] {A : Finset G} + +/-- A set with no doubling is either empty or the translate of a subgroup. + +Precisely, if `A` has no doubling then there exists a subgroup `H` such `aH = Ha = A` for all +`a ∈ A`. -/ +@[to_additive "A set with no doubling is either empty or the translate of a subgroup. + +Precisely, if `A` has no doubling then there exists a subgroup `H` such `a + H = H + a = A` for all +`a ∈ A`."] +lemma exists_subgroup_of_no_doubling (hA : #(A * A) ≤ #A) : + ∃ H : Subgroup G, ∀ a ∈ A, a •> (H : Set G) = A ∧ (H : Set G) <• a = A := by + have smul_A {a} (ha : a ∈ A) : a •> A = A * A := + eq_of_subset_of_card_le (smul_finset_subset_mul ha) (by simpa) + have A_smul {a} (ha : a ∈ A) : A <• a = A * A := + eq_of_subset_of_card_le (op_smul_finset_subset_mul ha) (by simpa) + have smul_A_eq_op_smul_A {a} (ha : a ∈ A) : a •> A = A <• a := by rw [smul_A ha, A_smul ha] + have smul_A_eq_op_smul_A' {a} (ha : a ∈ A) : a⁻¹ •> A = A <• a⁻¹ := by + rw [inv_smul_eq_iff, smul_comm, smul_A_eq_op_smul_A ha, op_inv, inv_smul_smul] + let H := stabilizer G A + have inv_smul_A {a} (ha : a ∈ A) : a⁻¹ • (A : Set G) = H := by + ext x + refine ⟨?_, fun hx ↦ ?_⟩ + · rintro ⟨b, hb, rfl⟩ + simp [H, mul_smul, inv_smul_eq_iff, smul_A ha, smul_A hb] + · norm_cast + rwa [smul_A_eq_op_smul_A' ha, op_inv, mem_inv_smul_finset_iff, op_smul_eq_mul, ← smul_eq_mul, + ← mem_inv_smul_finset_iff, inv_mem hx] + refine ⟨H, fun a ha ↦ ⟨?_, ?_⟩⟩ + · rw [← inv_smul_A ha, smul_inv_smul] + · rw [← inv_smul_A ha, smul_comm] + norm_cast + rw [← smul_A_eq_op_smul_A ha, inv_smul_smul] + +end Finset From 36ecbb9b41a8451306a4b8599cc45fdfdf8e3c27 Mon Sep 17 00:00:00 2001 From: Mitchell Lee Date: Tue, 10 Dec 2024 21:12:45 +0000 Subject: [PATCH 647/829] feat: infimum of filter coproduct and filter product (#19716) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We prove that if `f₁, f₂` are filters on a type α and `g₁, g₂` are filters on a type β, then `f₁.coprod g₁ ⊓ f₂ ×ˢ g₂ ≤ f₁ ×ˢ g₂ ⊔ f₂ ×ˢ g₁`. --- Mathlib/Order/Filter/Prod.lean | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/Mathlib/Order/Filter/Prod.lean b/Mathlib/Order/Filter/Prod.lean index 0f0c40960a7ab..652a0c81b4dc6 100644 --- a/Mathlib/Order/Filter/Prod.lean +++ b/Mathlib/Order/Filter/Prod.lean @@ -455,6 +455,15 @@ theorem coprod_neBot_left [NeBot f] [Nonempty β] : (f.coprod g).NeBot := theorem coprod_neBot_right [NeBot g] [Nonempty α] : (f.coprod g).NeBot := coprod_neBot_iff.2 (Or.inr ⟨‹_›, ‹_›⟩) +theorem coprod_inf_prod_le (f₁ f₂ : Filter α) (g₁ g₂ : Filter β) : + f₁.coprod g₁ ⊓ f₂ ×ˢ g₂ ≤ f₁ ×ˢ g₂ ⊔ f₂ ×ˢ g₁ := calc + f₁.coprod g₁ ⊓ f₂ ×ˢ g₂ + _ = (f₁ ×ˢ ⊤ ⊔ ⊤ ×ˢ g₁) ⊓ f₂ ×ˢ g₂ := by rw [coprod_eq_prod_top_sup_top_prod] + _ = f₁ ×ˢ ⊤ ⊓ f₂ ×ˢ g₂ ⊔ ⊤ ×ˢ g₁ ⊓ f₂ ×ˢ g₂ := inf_sup_right _ _ _ + _ = (f₁ ⊓ f₂) ×ˢ g₂ ⊔ f₂ ×ˢ (g₁ ⊓ g₂) := by simp [prod_inf_prod] + _ ≤ f₁ ×ˢ g₂ ⊔ f₂ ×ˢ g₁ := + sup_le_sup (prod_mono inf_le_left le_rfl) (prod_mono le_rfl inf_le_left) + theorem principal_coprod_principal (s : Set α) (t : Set β) : (𝓟 s).coprod (𝓟 t) = 𝓟 (sᶜ ×ˢ tᶜ)ᶜ := by rw [Filter.coprod, comap_principal, comap_principal, sup_principal, Set.prod_eq, compl_inter, From 75368bee46034cbb2e556461d176e4a41391cc9e Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Tue, 10 Dec 2024 21:48:39 +0000 Subject: [PATCH 648/829] chore(Data/Finsupp): split off basic scalar multiplication (#19205) In the comments of #19095 we discussed that `MonoidAlgebra` and `Polynomial` need scalar multiplication by `Nat` or `Int` early on, to define a `Semiring` or `Ring` structure. We can define a generic `SMulWithZero` instance and then specialize it to get `nsmul` and `zsmul`, but only if that instance is available early on for `Finsupp`. So: split this off from `Finsupp/Basic.lean` into a higher-up file. This PR used Damiano's `upstreamableDecls` linter to find out which results could move together with the definitions for free. --- Mathlib.lean | 1 + Mathlib/Data/Finsupp/Basic.lean | 56 +------------- Mathlib/Data/Finsupp/SMulWithZero.lean | 102 +++++++++++++++++++++++++ 3 files changed, 105 insertions(+), 54 deletions(-) create mode 100644 Mathlib/Data/Finsupp/SMulWithZero.lean diff --git a/Mathlib.lean b/Mathlib.lean index c5d560bd39014..2fdd9915d4878 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2498,6 +2498,7 @@ import Mathlib.Data.Finsupp.Notation import Mathlib.Data.Finsupp.Order import Mathlib.Data.Finsupp.PWO import Mathlib.Data.Finsupp.Pointwise +import Mathlib.Data.Finsupp.SMulWithZero import Mathlib.Data.Finsupp.ToDFinsupp import Mathlib.Data.Finsupp.Weight import Mathlib.Data.Finsupp.WellFounded diff --git a/Mathlib/Data/Finsupp/Basic.lean b/Mathlib/Data/Finsupp/Basic.lean index 6ba35e3f0ddf1..7c786c4ec0d7e 100644 --- a/Mathlib/Data/Finsupp/Basic.lean +++ b/Mathlib/Data/Finsupp/Basic.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.BigOperators.Finsupp import Mathlib.Algebra.Group.Action.Basic import Mathlib.Algebra.Module.Basic import Mathlib.Algebra.Regular.SMul +import Mathlib.Data.Finsupp.SMulWithZero import Mathlib.Data.Rat.BigOperators /-! @@ -1285,25 +1286,11 @@ end section -instance smulZeroClass [Zero M] [SMulZeroClass R M] : SMulZeroClass R (α →₀ M) where - smul a v := v.mapRange (a • ·) (smul_zero _) - smul_zero a := by - ext - apply smul_zero - /-! Throughout this section, some `Monoid` and `Semiring` arguments are specified with `{}` instead of `[]`. See note [implicit instance arguments]. -/ -@[simp, norm_cast] -theorem coe_smul [Zero M] [SMulZeroClass R M] (b : R) (v : α →₀ M) : ⇑(b • v) = b • ⇑v := - rfl - -theorem smul_apply [Zero M] [SMulZeroClass R M] (b : R) (v : α →₀ M) (a : α) : - (b • v) a = b • v a := - rfl - theorem _root_.IsSMulRegular.finsupp [Zero M] [SMulZeroClass R M] {k : R} (hk : IsSMulRegular M k) : IsSMulRegular (α →₀ M) k := fun _ _ h => ext fun i => hk (DFunLike.congr_fun h i) @@ -1314,34 +1301,14 @@ instance faithfulSMul [Nonempty α] [Zero M] [SMulZeroClass R M] [FaithfulSMul R let ⟨a⟩ := ‹Nonempty α› eq_of_smul_eq_smul fun m : M => by simpa using DFunLike.congr_fun (h (single a m)) a -instance instSMulWithZero [Zero R] [Zero M] [SMulWithZero R M] : SMulWithZero R (α →₀ M) where - zero_smul f := by ext i; exact zero_smul _ _ - variable (α M) -instance distribSMul [AddZeroClass M] [DistribSMul R M] : DistribSMul R (α →₀ M) where - smul := (· • ·) - smul_add _ _ _ := ext fun _ => smul_add _ _ _ - smul_zero _ := ext fun _ => smul_zero _ - instance distribMulAction [Monoid R] [AddMonoid M] [DistribMulAction R M] : DistribMulAction R (α →₀ M) := { Finsupp.distribSMul _ _ with one_smul := fun x => ext fun y => one_smul R (x y) mul_smul := fun r s x => ext fun y => mul_smul r s (x y) } -instance isScalarTower [Zero M] [SMulZeroClass R M] [SMulZeroClass S M] [SMul R S] - [IsScalarTower R S M] : IsScalarTower R S (α →₀ M) where - smul_assoc _ _ _ := ext fun _ => smul_assoc _ _ _ - -instance smulCommClass [Zero M] [SMulZeroClass R M] [SMulZeroClass S M] [SMulCommClass R S M] : - SMulCommClass R S (α →₀ M) where - smul_comm _ _ _ := ext fun _ => smul_comm _ _ _ - -instance isCentralScalar [Zero M] [SMulZeroClass R M] [SMulZeroClass Rᵐᵒᵖ M] [IsCentralScalar R M] : - IsCentralScalar R (α →₀ M) where - op_smul_eq_smul _ _ := ext fun _ => op_smul_eq_smul _ _ - instance module [Semiring R] [AddCommMonoid M] [Module R M] : Module R (α →₀ M) := { toDistribMulAction := Finsupp.distribMulAction α M zero_smul := fun _ => ext fun _ => zero_smul _ _ @@ -1349,11 +1316,6 @@ instance module [Semiring R] [AddCommMonoid M] [Module R M] : Module R (α → variable {α M} -theorem support_smul [AddMonoid M] [SMulZeroClass R M] {b : R} {g : α →₀ M} : - (b • g).support ⊆ g.support := fun a => by - simp only [smul_apply, mem_support_iff, Ne] - exact mt fun h => h.symm ▸ smul_zero _ - @[simp] theorem support_smul_eq [Semiring R] [AddCommMonoid M] [Module R M] [NoZeroSMulDivisors R M] {b : R} (hb : b ≠ 0) {g : α →₀ M} : (b • g).support = g.support := @@ -1376,25 +1338,11 @@ theorem mapDomain_smul {_ : Monoid R} [AddCommMonoid M] [DistribMulAction R M] { (v : α →₀ M) : mapDomain f (b • v) = b • mapDomain f v := mapDomain_mapRange _ _ _ _ (smul_add b) -@[simp] -theorem smul_single [Zero M] [SMulZeroClass R M] (c : R) (a : α) (b : M) : - c • Finsupp.single a b = Finsupp.single a (c • b) := - mapRange_single - -- Porting note: removed `simp` because `simpNF` can prove it. theorem smul_single' {_ : Semiring R} (c : R) (a : α) (b : R) : c • Finsupp.single a b = Finsupp.single a (c * b) := smul_single _ _ _ -theorem mapRange_smul {_ : Monoid R} [AddMonoid M] [DistribMulAction R M] [AddMonoid N] - [DistribMulAction R N] {f : M → N} {hf : f 0 = 0} (c : R) (v : α →₀ M) - (hsmul : ∀ x, f (c • x) = c • f x) : mapRange f hf (c • v) = c • mapRange f hf v := by - erw [← mapRange_comp] - · have : f ∘ (c • ·) = (c • ·) ∘ f := funext hsmul - simp_rw [this] - apply mapRange_comp - simp only [Function.comp_apply, smul_zero, hf] - theorem smul_single_one [Semiring R] (a : α) (b : R) : b • single a (1 : R) = single a b := by rw [smul_single, smul_eq_mul, mul_one] @@ -1701,4 +1649,4 @@ end Sigma end Finsupp -set_option linter.style.longFile 1900 +set_option linter.style.longFile 1700 diff --git a/Mathlib/Data/Finsupp/SMulWithZero.lean b/Mathlib/Data/Finsupp/SMulWithZero.lean new file mode 100644 index 0000000000000..e7e1adc5d920c --- /dev/null +++ b/Mathlib/Data/Finsupp/SMulWithZero.lean @@ -0,0 +1,102 @@ +/- +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, Kim Morrison +-/ +import Mathlib.Algebra.Group.Action.Pi +import Mathlib.Algebra.SMulWithZero +import Mathlib.Data.Finsupp.Defs + +/-! +# Scalar multiplication on `Finsupp` + +This file defines the pointwise scalar multiplication on `Finsupp`, assuming it preserves zero. + +## Main declarations + +* `Finsupp.smulZeroClass`: if the action of `R` on `M` preserves `0`, then it acts on `α →₀ M` + +## Implementation notes + +This file is intermediate between `Finsupp.Defs` and `Finsupp.Module` in that it covers scalar +multiplication but does not rely on the definition of `Module`. Scalar multiplication is needed to +supply the `nsmul` (and `zsmul`) fields of (semi)ring structures which are fundamental for e.g. +`Polynomial`, so we want to keep the imports requied for the `Finsupp.smulZeroClass` instance +reasonably light. + +This file is a `noncomputable theory` and uses classical logic throughout. +-/ + +assert_not_exists Module + +noncomputable section + +open Finset Function + +variable {α β γ ι M M' N P G H R S : Type*} + +namespace Finsupp + +instance smulZeroClass [Zero M] [SMulZeroClass R M] : SMulZeroClass R (α →₀ M) where + smul a v := v.mapRange (a • ·) (smul_zero _) + smul_zero a := by + ext + apply smul_zero + +/-! +Throughout this section, some `Monoid` and `Semiring` arguments are specified with `{}` instead of +`[]`. See note [implicit instance arguments]. +-/ + +@[simp, norm_cast] +theorem coe_smul [Zero M] [SMulZeroClass R M] (b : R) (v : α →₀ M) : ⇑(b • v) = b • ⇑v := + rfl + +theorem smul_apply [Zero M] [SMulZeroClass R M] (b : R) (v : α →₀ M) (a : α) : + (b • v) a = b • v a := + rfl + +instance instSMulWithZero [Zero R] [Zero M] [SMulWithZero R M] : SMulWithZero R (α →₀ M) where + zero_smul f := by ext i; exact zero_smul _ _ + +variable (α M) + +instance distribSMul [AddZeroClass M] [DistribSMul R M] : DistribSMul R (α →₀ M) where + smul := (· • ·) + smul_add _ _ _ := ext fun _ => smul_add _ _ _ + smul_zero _ := ext fun _ => smul_zero _ + +instance isScalarTower [Zero M] [SMulZeroClass R M] [SMulZeroClass S M] [SMul R S] + [IsScalarTower R S M] : IsScalarTower R S (α →₀ M) where + smul_assoc _ _ _ := ext fun _ => smul_assoc _ _ _ + +instance smulCommClass [Zero M] [SMulZeroClass R M] [SMulZeroClass S M] [SMulCommClass R S M] : + SMulCommClass R S (α →₀ M) where + smul_comm _ _ _ := ext fun _ => smul_comm _ _ _ + +instance isCentralScalar [Zero M] [SMulZeroClass R M] [SMulZeroClass Rᵐᵒᵖ M] [IsCentralScalar R M] : + IsCentralScalar R (α →₀ M) where + op_smul_eq_smul _ _ := ext fun _ => op_smul_eq_smul _ _ + +variable {α M} + +theorem support_smul [AddMonoid M] [SMulZeroClass R M] {b : R} {g : α →₀ M} : + (b • g).support ⊆ g.support := fun a => by + simp only [smul_apply, mem_support_iff, Ne] + exact mt fun h => h.symm ▸ smul_zero _ + +@[simp] +theorem smul_single [Zero M] [SMulZeroClass R M] (c : R) (a : α) (b : M) : + c • Finsupp.single a b = Finsupp.single a (c • b) := + mapRange_single + +theorem mapRange_smul {_ : Monoid R} [AddMonoid M] [DistribMulAction R M] [AddMonoid N] + [DistribMulAction R N] {f : M → N} {hf : f 0 = 0} (c : R) (v : α →₀ M) + (hsmul : ∀ x, f (c • x) = c • f x) : mapRange f hf (c • v) = c • mapRange f hf v := by + erw [← mapRange_comp] + · have : f ∘ (c • ·) = (c • ·) ∘ f := funext hsmul + simp_rw [this] + apply mapRange_comp + simp only [Function.comp_apply, smul_zero, hf] + +end Finsupp From 7bdc7fa6e695707d4ee2bc8678437d65063827d8 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 11 Dec 2024 00:00:34 +0000 Subject: [PATCH 649/829] chore: fix 'lake exe pole' and benchmark summary bot (#19878) --- .github/workflows/nightly_detect_failure.yml | 2 +- LongestPole/Main.lean | 6 ++++-- LongestPole/SpeedCenterJson.lean | 4 ++-- LongestPole/Unused.lean | 2 ++ scripts/bench_summary.lean | 4 ++-- 5 files changed, 11 insertions(+), 7 deletions(-) diff --git a/.github/workflows/nightly_detect_failure.yml b/.github/workflows/nightly_detect_failure.yml index 0f52fa1c48fc2..b37d8c9217e53 100644 --- a/.github/workflows/nightly_detect_failure.yml +++ b/.github/workflows/nightly_detect_failure.yml @@ -51,7 +51,7 @@ jobs: git tag "nightly-testing-${version}" git push origin "nightly-testing-${version}" hash="$(git rev-parse "nightly-testing-${version}")" - curl -X POST "http://speed.lean-fro.org/mathlib4/api/queue/commit/e7b27246-a3e6-496a-b552-ff4b45c7236e/$hash" -u "admin:${{ secrets.SPEED }}" + curl -X POST "https://speed.lean-lang.org/mathlib4/api/queue/commit/e7b27246-a3e6-496a-b552-ff4b45c7236e/$hash" -u "admin:${{ secrets.SPEED }}" fi hash="$(git rev-parse "nightly-testing-${version}")" printf 'SHA=%s\n' "${hash}" >> "${GITHUB_ENV}" diff --git a/LongestPole/Main.lean b/LongestPole/Main.lean index 0ab7b0d3148a6..102f398c4acf7 100644 --- a/LongestPole/Main.lean +++ b/LongestPole/Main.lean @@ -34,7 +34,7 @@ def mathlib4RepoId : String := "e7b27246-a3e6-496a-b552-ff4b45c7236e" namespace SpeedCenterAPI def runJson (hash : String) (repoId : String := mathlib4RepoId) : IO String := - runCurl #[s!"http://speed.lean-fro.org/mathlib4/api/run/{repoId}?hash={hash}"] + runCurl #[s!"https://speed.lean-lang.org/mathlib4/api/run/{repoId}?hash={hash}"] def getRunResponse (hash : String) : IO RunResponse := do let r ← runJson hash @@ -44,7 +44,7 @@ def getRunResponse (hash : String) : IO RunResponse := do | .ok v => pure v | .error e => match fromJson? j with | .ok (v : ErrorMessage) => - IO.eprintln s!"http://speed.lean-fro.org says: {v.message}" + IO.eprintln s!"https://speed.lean-lang.org says: {v.message}" IO.eprintln s!"Try moving to an older commit?" IO.Process.exit 1 | .error _ => throw <| IO.userError s!"Could not parse speed center JSON: {e}\n{j}" @@ -135,6 +135,8 @@ def longestPoleCLI (args : Cli.Parsed) : IO UInt32 := do | some to => pure <| to.as! ModuleName | none => ImportGraph.getCurrentModule -- autodetect the main module from the `lakefile.lean` searchPathRef.set compile_time_search_path% + -- It may be reasonable to remove this again after https://github.com/leanprover/lean4/pull/6325 + unsafe enableInitializersExecution unsafe withImportModules #[{module := to}] {} (trustLevel := 1024) fun env => do let graph := env.importGraph let sha ← headSha diff --git a/LongestPole/SpeedCenterJson.lean b/LongestPole/SpeedCenterJson.lean index 20a39f872dd0c..007d464cc463a 100644 --- a/LongestPole/SpeedCenterJson.lean +++ b/LongestPole/SpeedCenterJson.lean @@ -7,7 +7,7 @@ import Lean.Data.Json open Lean /-! -# `structure`s for the http://speed.lean-fro.org/ API +# `structure`s for the https://speed.lean-lang.org/ API -/ namespace SpeedCenterAPI @@ -42,7 +42,7 @@ structure Run where result : Result deriving ToJson, FromJson -/-- The top-level API response for `http://speed.lean-fro.org/{repo}/api/run/{guid}?hash={sha}`. -/ +/-- The top-level API response for `https://speed.lean-lang.org/{repo}/api/run/{guid}?hash={sha}`. -/ structure RunResponse where run : Run deriving ToJson, FromJson diff --git a/LongestPole/Unused.lean b/LongestPole/Unused.lean index ebf4d74987ec3..5e405339d7b4b 100644 --- a/LongestPole/Unused.lean +++ b/LongestPole/Unused.lean @@ -63,6 +63,8 @@ def unusedImportsCLI (args : Cli.Parsed) : IO UInt32 := do -- The code below assumes that it is "deeper files first", as reported by `lake exe pole`. searchPathRef.set compile_time_search_path% + -- It may be reasonable to remove this again after https://github.com/leanprover/lean4/pull/6325 + unsafe enableInitializersExecution let (unused, _) ← unsafe withImportModules #[{module := `Mathlib}] {} (trustLevel := 1024) fun env => Prod.fst <$> Core.CoreM.toIO (ctx := { fileName := "", fileMap := default }) (s := { env := env }) do diff --git a/scripts/bench_summary.lean b/scripts/bench_summary.lean index 17f0467b2ef97..9491ddef2ee99 100644 --- a/scripts/bench_summary.lean +++ b/scripts/bench_summary.lean @@ -192,9 +192,9 @@ def addBenchSummaryComment (PR : Nat) (repo : String) (jobID : String := "") -- retrieve the data from the speed-center let curlSpeedCenter : IO.Process.SpawnArgs := { cmd := "curl" - args := #[s!"http://speed.lean-fro.org/mathlib4/api/compare/{source}/to/{target}?all_values=true"] } + args := #[s!"https://speed.lean-lang.org/mathlib4/api/compare/{source}/to/{target}?all_values=true"] } dbg_trace "\n#running\n\ - curl http://speed.lean-fro.org/mathlib4/api/compare/{source}/to/{target}?all_values=true > {tempFile}.src" + curl https://speed.lean-lang.org/mathlib4/api/compare/{source}/to/{target}?all_values=true > {tempFile}.src" let bench ← IO.Process.run curlSpeedCenter IO.FS.writeFile (tempFile ++ ".src") bench From 43404d79f195610c4184a02f530bdc36c5b2285a Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 11 Dec 2024 04:11:28 +0000 Subject: [PATCH 650/829] chore: used named arguments instead of @DecidableRel (#19850) --- .../Order/GroupWithZero/Unbundled.lean | 2 +- Mathlib/Algebra/Order/Ring/Defs.lean | 10 +++--- Mathlib/Combinatorics/Colex.lean | 6 ++-- .../Enumerative/IncidenceAlgebra.lean | 26 +++++++------- .../SetFamily/AhlswedeZhang.lean | 18 +++++----- .../SetFamily/KruskalKatona.lean | 3 +- Mathlib/Data/DFinsupp/Lex.lean | 4 +-- Mathlib/Data/Finset/Defs.lean | 8 ++--- Mathlib/Data/Finset/Sups.lean | 9 ++--- Mathlib/Data/List/MinMax.lean | 4 +-- Mathlib/Data/Num/Basic.lean | 12 +++---- Mathlib/Data/Ordmap/Ordnode.lean | 6 ++-- Mathlib/Data/Ordmap/Ordset.lean | 34 +++++++++---------- Mathlib/Data/String/Basic.lean | 4 +-- Mathlib/Logic/Encodable/Basic.lean | 2 +- Mathlib/NumberTheory/Zsqrtd/Basic.lean | 2 +- Mathlib/Order/Antisymmetrization.lean | 2 +- Mathlib/Order/Basic.lean | 13 +++---- Mathlib/Order/Chain.lean | 2 +- Mathlib/Order/Compare.lean | 19 ++++++----- Mathlib/Order/Defs/PartialOrder.lean | 6 ++-- Mathlib/Order/Interval/Basic.lean | 10 +++--- Mathlib/Order/Interval/Finset/Box.lean | 2 +- Mathlib/Order/Interval/Finset/Defs.lean | 12 +++---- Mathlib/Order/SuccPred/Archimedean.lean | 8 ++--- Mathlib/Order/WithBot.lean | 12 +++---- MathlibTest/instance_diamonds.lean | 4 +-- 27 files changed, 125 insertions(+), 115 deletions(-) diff --git a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean index dc2a61ca1a817..01454ee338502 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean @@ -1107,7 +1107,7 @@ lemma strictMonoOn_mul_self [PosMulStrictMono M₀] [MulPosMono M₀] : -- See Note [decidable namespace] protected lemma Decidable.mul_lt_mul'' [PosMulMono M₀] [PosMulStrictMono M₀] [MulPosStrictMono M₀] - [@DecidableRel M₀ (· ≤ ·)] (h1 : a < c) (h2 : b < d) + [DecidableRel (α := M₀) (· ≤ ·)] (h1 : a < c) (h2 : b < d) (h3 : 0 ≤ a) (h4 : 0 ≤ b) : a * b < c * d := h4.lt_or_eq_dec.elim (fun b0 ↦ mul_lt_mul h1 h2.le b0 <| h3.trans h1.le) fun b0 ↦ by rw [← b0, mul_zero]; exact mul_pos (h3.trans_lt h1) (h4.trans_lt h2) diff --git a/Mathlib/Algebra/Order/Ring/Defs.lean b/Mathlib/Algebra/Order/Ring/Defs.lean index d9128b14a9a37..a8f2ee51df52a 100644 --- a/Mathlib/Algebra/Order/Ring/Defs.lean +++ b/Mathlib/Algebra/Order/Ring/Defs.lean @@ -265,7 +265,8 @@ instance (priority := 200) StrictOrderedSemiring.toMulPosStrictMono : MulPosStri -- See note [reducible non-instances] /-- A choice-free version of `StrictOrderedSemiring.toOrderedSemiring` to avoid using choice in basic `Nat` lemmas. -/ -abbrev StrictOrderedSemiring.toOrderedSemiring' [@DecidableRel α (· ≤ ·)] : OrderedSemiring α := +abbrev StrictOrderedSemiring.toOrderedSemiring' [DecidableRel (α := α) (· ≤ ·)] : + OrderedSemiring α := { ‹StrictOrderedSemiring α› with mul_le_mul_of_nonneg_left := fun a b c hab hc => by obtain rfl | hab := Decidable.eq_or_lt_of_le hab @@ -308,7 +309,7 @@ variable [StrictOrderedCommSemiring α] -- See note [reducible non-instances] /-- A choice-free version of `StrictOrderedCommSemiring.toOrderedCommSemiring'` to avoid using choice in basic `Nat` lemmas. -/ -abbrev StrictOrderedCommSemiring.toOrderedCommSemiring' [@DecidableRel α (· ≤ ·)] : +abbrev StrictOrderedCommSemiring.toOrderedCommSemiring' [DecidableRel (α := α) (· ≤ ·)] : OrderedCommSemiring α := { ‹StrictOrderedCommSemiring α›, StrictOrderedSemiring.toOrderedSemiring' with } @@ -334,7 +335,7 @@ instance (priority := 100) StrictOrderedRing.toStrictOrderedSemiring : StrictOrd -- See note [reducible non-instances] /-- A choice-free version of `StrictOrderedRing.toOrderedRing` to avoid using choice in basic `Int` lemmas. -/ -abbrev StrictOrderedRing.toOrderedRing' [@DecidableRel α (· ≤ ·)] : OrderedRing α := +abbrev StrictOrderedRing.toOrderedRing' [DecidableRel (α := α) (· ≤ ·)] : OrderedRing α := { ‹StrictOrderedRing α›, (Ring.toSemiring : Semiring α) with mul_nonneg := fun a b ha hb => by obtain ha | ha := Decidable.eq_or_lt_of_le ha @@ -357,7 +358,8 @@ variable [StrictOrderedCommRing α] -- See note [reducible non-instances] /-- A choice-free version of `StrictOrderedCommRing.toOrderedCommRing` to avoid using choice in basic `Int` lemmas. -/ -abbrev StrictOrderedCommRing.toOrderedCommRing' [@DecidableRel α (· ≤ ·)] : OrderedCommRing α := +abbrev StrictOrderedCommRing.toOrderedCommRing' [DecidableRel (α := α) (· ≤ ·)] : + OrderedCommRing α := { ‹StrictOrderedCommRing α›, StrictOrderedRing.toOrderedRing' with } -- See note [lower instance priority] diff --git a/Mathlib/Combinatorics/Colex.lean b/Mathlib/Combinatorics/Colex.lean index 7c732db7e7473..98e0a7d8d89ae 100644 --- a/Mathlib/Combinatorics/Colex.lean +++ b/Mathlib/Combinatorics/Colex.lean @@ -220,11 +220,11 @@ variable [DecidableEq α] instance instDecidableEq : DecidableEq (Colex α) := fun s t ↦ decidable_of_iff' (s.ofColex = t.ofColex) Colex.ext_iff -instance instDecidableLE [@DecidableRel α (· ≤ ·)] : @DecidableRel (Colex α) (· ≤ ·) := fun s t ↦ - decidable_of_iff' +instance instDecidableLE [DecidableRel (α := α) (· ≤ ·)] : DecidableRel (α := Colex α) (· ≤ ·) := + fun s t ↦ decidable_of_iff' (∀ ⦃a⦄, a ∈ ofColex s → a ∉ ofColex t → ∃ b, b ∈ ofColex t ∧ b ∉ ofColex s ∧ a ≤ b) Iff.rfl -instance instDecidableLT [@DecidableRel α (· ≤ ·)] : @DecidableRel (Colex α) (· < ·) := +instance instDecidableLT [DecidableRel (α := α) (· ≤ ·)] : DecidableRel (α := Colex α) (· < ·) := decidableLTOfDecidableLE /-- The colexigraphic order is insensitive to removing the same elements from both sets. -/ diff --git a/Mathlib/Combinatorics/Enumerative/IncidenceAlgebra.lean b/Mathlib/Combinatorics/Enumerative/IncidenceAlgebra.lean index a049787977d46..cb22ef58fde04 100644 --- a/Mathlib/Combinatorics/Enumerative/IncidenceAlgebra.lean +++ b/Mathlib/Combinatorics/Enumerative/IncidenceAlgebra.lean @@ -321,7 +321,7 @@ instance algebraRight [PartialOrder α] [LocallyFiniteOrder α] [DecidableEq α] /-! ### The Lambda function -/ section Lambda -variable (𝕜) [Zero 𝕜] [One 𝕜] [Preorder α] [@DecidableRel α (· ⩿ ·)] +variable (𝕜) [Zero 𝕜] [One 𝕜] [Preorder α] [DecidableRel (α := α) (· ⩿ ·)] /-- The lambda function of the incidence algebra is the function that assigns `1` to every nonempty interval of cardinality one or two. -/ @@ -334,7 +334,7 @@ end Lambda /-! ### The Zeta and Möbius functions -/ section Zeta -variable (𝕜) [Zero 𝕜] [One 𝕜] [LE α] [@DecidableRel α (· ≤ ·)] {a b : α} +variable (𝕜) [Zero 𝕜] [One 𝕜] [LE α] [DecidableRel (α := α) (· ≤ ·)] {a b : α} /-- The zeta function of the incidence algebra is the function that assigns 1 to every nonempty interval, convolution with this function sums functions over intervals. -/ @@ -348,15 +348,16 @@ lemma zeta_of_le (h : a ≤ b) : zeta 𝕜 a b = 1 := if_pos h end Zeta -lemma zeta_mul_zeta [Semiring 𝕜] [Preorder α] [LocallyFiniteOrder α] [@DecidableRel α (· ≤ ·)] +lemma zeta_mul_zeta [Semiring 𝕜] [Preorder α] [LocallyFiniteOrder α] [DecidableRel (α := α) (· ≤ ·)] (a b : α) : (zeta 𝕜 * zeta 𝕜 : IncidenceAlgebra 𝕜 α) a b = (Icc a b).card := by rw [mul_apply, card_eq_sum_ones, Nat.cast_sum, Nat.cast_one] refine sum_congr rfl fun x hx ↦ ?_ rw [mem_Icc] at hx rw [zeta_of_le hx.1, zeta_of_le hx.2, one_mul] -lemma zeta_mul_kappa [Semiring 𝕜] [Preorder α] [LocallyFiniteOrder α] [@DecidableRel α (· ≤ ·)] - (a b : α) : (zeta 𝕜 * zeta 𝕜 : IncidenceAlgebra 𝕜 α) a b = (Icc a b).card := by +lemma zeta_mul_kappa [Semiring 𝕜] [Preorder α] [LocallyFiniteOrder α] + [DecidableRel (α := α) (· ≤ ·)] (a b : α) : + (zeta 𝕜 * zeta 𝕜 : IncidenceAlgebra 𝕜 α) a b = (Icc a b).card := by rw [mul_apply, card_eq_sum_ones, Nat.cast_sum, Nat.cast_one] refine sum_congr rfl fun x hx ↦ ?_ rw [mem_Icc] at hx @@ -474,7 +475,7 @@ end Mu'Spec section MuZeta variable (𝕜 α) [AddCommGroup 𝕜] [MulOneClass 𝕜] [PartialOrder α] [LocallyFiniteOrder α] - [DecidableEq α] [@DecidableRel α (· ≤ ·)] + [DecidableEq α] [DecidableRel (α := α) (· ≤ ·)] lemma mu_mul_zeta : (mu 𝕜 * zeta 𝕜 : IncidenceAlgebra 𝕜 α) = 1 := by ext a b @@ -500,7 +501,7 @@ private lemma mu_eq_mu' : (mu 𝕜 : IncidenceAlgebra 𝕜 α) = mu' 𝕜 := by lemma mu_eq_neg_sum_Ioc_of_ne (hab : a ≠ b) : mu 𝕜 a b = -∑ x ∈ Ioc a b, mu 𝕜 x b := by rw [mu_eq_mu', mu'_eq_sum_Ioc_of_ne hab] -lemma zeta_mul_mu [@DecidableRel α (· ≤ ·)] : (zeta 𝕜 * mu 𝕜 : IncidenceAlgebra 𝕜 α) = 1 := by +lemma zeta_mul_mu [DecidableRel (α := α) (· ≤ ·)] : (zeta 𝕜 * mu 𝕜 : IncidenceAlgebra 𝕜 α) = 1 := by rw [mu_eq_mu', zeta_mul_mu'] lemma sum_Icc_mu_left (a b : α) : ∑ x ∈ Icc a b, mu 𝕜 x b = if a = b then 1 else 0 := by @@ -513,7 +514,7 @@ variable (𝕜) [Ring 𝕜] [PartialOrder α] [LocallyFiniteOrder α] [Decidable @[simp] lemma mu_toDual (a b : α) : mu 𝕜 (toDual a) (toDual b) = mu 𝕜 b a := by - letI : @DecidableRel α (· ≤ ·) := Classical.decRel _ + letI : DecidableRel (α := α) (· ≤ ·) := Classical.decRel _ let mud : IncidenceAlgebra 𝕜 αᵒᵈ := { toFun := fun a b ↦ mu 𝕜 (ofDual b) (ofDual a) eq_zero_of_not_le' := fun a b hab ↦ apply_eq_zero_of_not_le (by exact hab) _ } @@ -544,7 +545,7 @@ variable [Ring 𝕜] [PartialOrder α] [OrderTop α] [LocallyFiniteOrder α] [De O'Donnell. -/ lemma moebius_inversion_top (f g : α → 𝕜) (h : ∀ x, g x = ∑ y ∈ Ici x, f y) (x : α) : f x = ∑ y ∈ Ici x, mu 𝕜 x y * g y := by - letI : @DecidableRel α (· ≤ ·) := Classical.decRel _ + letI : DecidableRel (α := α) (· ≤ ·) := Classical.decRel _ symm calc ∑ y ∈ Ici x, mu 𝕜 x y * g y = ∑ y ∈ Ici x, mu 𝕜 x y * ∑ z ∈ Ici y, f z := by simp_rw [h] @@ -612,7 +613,8 @@ lemma prod_mk (a₁ a₂ : α) (b₁ b₂ : β) : f.prod g (a₁, b₁) (a₂, b @[simp] lemma prod_apply (x y : α × β) : f.prod g x y = f x.1 y.1 * g x.2 y.2 := rfl /-- This is a version of `IncidenceAlgebra.prod_mul_prod` that works over non-commutative rings. -/ -lemma prod_mul_prod' [LocallyFiniteOrder α] [LocallyFiniteOrder β] [@DecidableRel (α × β) (· ≤ ·)] +lemma prod_mul_prod' [LocallyFiniteOrder α] [LocallyFiniteOrder β] + [DecidableRel (α := α × β) (· ≤ ·)] (h : ∀ a₁ a₂ a₃ b₁ b₂ b₃, f₁ a₁ a₂ * g₁ b₁ b₂ * (f₂ a₂ a₃ * g₂ b₂ b₃) = f₁ a₁ a₂ * f₂ a₂ a₃ * (g₁ b₁ b₂ * g₂ b₂ b₃)) : f₁.prod g₁ * f₂.prod g₂ = (f₁ * f₂).prod (g₁ * g₂) := by @@ -624,7 +626,7 @@ lemma one_prod_one [DecidableEq α] [DecidableEq β] : ext x y; simp [Prod.ext_iff, ← ite_and, and_comm] @[simp] -lemma zeta_prod_zeta [@DecidableRel α (· ≤ ·)] [@DecidableRel β (· ≤ ·)] : +lemma zeta_prod_zeta [DecidableRel (α := α) (· ≤ ·)] [DecidableRel (α := β) (· ≤ ·)] : (zeta 𝕜).prod (zeta 𝕜) = (zeta 𝕜 : IncidenceAlgebra 𝕜 (α × β)) := by ext x y hxy; simp [hxy, hxy.1, hxy.2] @@ -632,7 +634,7 @@ end Ring section CommRing variable [CommRing 𝕜] [Preorder α] [Preorder β] [LocallyFiniteOrder α] [LocallyFiniteOrder β] - [@DecidableRel (α × β) (· ≤ ·)] (f₁ f₂ : IncidenceAlgebra 𝕜 α) (g₁ g₂ : IncidenceAlgebra 𝕜 β) + [DecidableRel (α := α × β) (· ≤ ·)] (f₁ f₂ : IncidenceAlgebra 𝕜 α) (g₁ g₂ : IncidenceAlgebra 𝕜 β) @[simp] lemma prod_mul_prod : f₁.prod g₁ * f₂.prod g₂ = (f₁ * f₂).prod (g₁ * g₂) := diff --git a/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean b/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean index be6ce35efc01e..aaffbd0028f58 100644 --- a/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean +++ b/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean @@ -103,14 +103,15 @@ variable {α β : Type*} section SemilatticeSup variable [SemilatticeSup α] [SemilatticeSup β] [BoundedOrder β] {s t : Finset α} {a : α} -private lemma sup_aux [@DecidableRel α (· ≤ ·)] : a ∈ lowerClosure s → {b ∈ s | a ≤ b}.Nonempty := +private lemma sup_aux [DecidableRel (α := α) (· ≤ ·)] : + a ∈ lowerClosure s → {b ∈ s | a ≤ b}.Nonempty := fun ⟨b, hb, hab⟩ ↦ ⟨b, mem_filter.2 ⟨hb, hab⟩⟩ private lemma lower_aux [DecidableEq α] : a ∈ lowerClosure ↑(s ∪ t) ↔ a ∈ lowerClosure s ∨ a ∈ lowerClosure t := by rw [coe_union, lowerClosure_union, LowerSet.mem_sup_iff] -variable [@DecidableRel α (· ≤ ·)] [OrderTop α] +variable [DecidableRel (α := α) (· ≤ ·)] [OrderTop α] /-- The supremum of the elements of `s` less than `a` if there are some, otherwise `⊤`. -/ def truncatedSup (s : Finset α) (a : α) : α := @@ -133,7 +134,7 @@ lemma le_truncatedSup : a ≤ truncatedSup s a := by exact h.trans <| le_sup' id <| mem_filter.2 ⟨hb, h⟩ · exact le_top -lemma map_truncatedSup [@DecidableRel β (· ≤ ·)] (e : α ≃o β) (s : Finset α) (a : α) : +lemma map_truncatedSup [DecidableRel (α := β) (· ≤ ·)] (e : α ≃o β) (s : Finset α) (a : α) : e (truncatedSup s a) = truncatedSup (s.map e.toEquiv.toEmbedding) (e a) := by have : e a ∈ lowerClosure (s.map e.toEquiv.toEmbedding : Set β) ↔ a ∈ lowerClosure s := by simp simp_rw [truncatedSup, apply_dite e, map_finset_sup', map_top, this] @@ -174,16 +175,17 @@ end SemilatticeSup section SemilatticeInf variable [SemilatticeInf α] [SemilatticeInf β] - [BoundedOrder β] [@DecidableRel β (· ≤ ·)] {s t : Finset α} {a : α} + [BoundedOrder β] [DecidableRel (α := β) (· ≤ ·)] {s t : Finset α} {a : α} -private lemma inf_aux [@DecidableRel α (· ≤ ·)]: a ∈ upperClosure s → {b ∈ s | b ≤ a}.Nonempty := +private lemma inf_aux [DecidableRel (α := α) (· ≤ ·)] : + a ∈ upperClosure s → {b ∈ s | b ≤ a}.Nonempty := fun ⟨b, hb, hab⟩ ↦ ⟨b, mem_filter.2 ⟨hb, hab⟩⟩ private lemma upper_aux [DecidableEq α] : a ∈ upperClosure ↑(s ∪ t) ↔ a ∈ upperClosure s ∨ a ∈ upperClosure t := by rw [coe_union, upperClosure_union, UpperSet.mem_inf_iff] -variable [@DecidableRel α (· ≤ ·)] [BoundedOrder α] +variable [DecidableRel (α := α) (· ≤ ·)] [BoundedOrder α] /-- The infimum of the elements of `s` less than `a` if there are some, otherwise `⊥`. -/ def truncatedInf (s : Finset α) (a : α) : α := @@ -256,7 +258,7 @@ private lemma infs_aux : a ∈ lowerClosure ↑(s ⊼ t) ↔ a ∈ lowerClosure private lemma sups_aux : a ∈ upperClosure ↑(s ⊻ t) ↔ a ∈ upperClosure s ∧ a ∈ upperClosure t := by rw [coe_sups, upperClosure_sups, UpperSet.mem_sup_iff] -variable [@DecidableRel α (· ≤ ·)] [BoundedOrder α] +variable [DecidableRel (α := α) (· ≤ ·)] [BoundedOrder α] lemma truncatedSup_infs (hs : a ∈ lowerClosure s) (ht : a ∈ lowerClosure t) : truncatedSup (s ⊼ t) a = truncatedSup s a ⊓ truncatedSup t a := by @@ -283,7 +285,7 @@ lemma truncatedInf_sups_of_not_mem (ha : a ∉ upperClosure s ⊔ upperClosure t end DistribLattice section BooleanAlgebra -variable [BooleanAlgebra α] [@DecidableRel α (· ≤ ·)] +variable [BooleanAlgebra α] [DecidableRel (α := α) (· ≤ ·)] @[simp] lemma compl_truncatedSup (s : Finset α) (a : α) : (truncatedSup s a)ᶜ = truncatedInf sᶜˢ aᶜ := map_truncatedSup (OrderIso.compl α) _ _ diff --git a/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean b/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean index edfdf5d218350..d3655a022cb05 100644 --- a/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean +++ b/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean @@ -143,7 +143,8 @@ lemma toColex_compress_lt_toColex {hU : U.Nonempty} {hV : V.Nonempty} (h : max' private def UsefulCompression (U V : Finset α) : Prop := Disjoint U V ∧ #U = #V ∧ ∃ (HU : U.Nonempty) (HV : V.Nonempty), max' U HU < max' V HV -private instance UsefulCompression.instDecidableRel : @DecidableRel (Finset α) UsefulCompression := +private instance UsefulCompression.instDecidableRel : + DecidableRel (α := Finset α) UsefulCompression := fun _ _ ↦ inferInstanceAs (Decidable (_ ∧ _)) /-- Applying a good compression will decrease measure, keep cardinality, keep sizes and decrease diff --git a/Mathlib/Data/DFinsupp/Lex.lean b/Mathlib/Data/DFinsupp/Lex.lean index 1b3888f712e29..3251cba793ca2 100644 --- a/Mathlib/Data/DFinsupp/Lex.lean +++ b/Mathlib/Data/DFinsupp/Lex.lean @@ -91,13 +91,13 @@ private def lt_trichotomy_rec {P : Lex (Π₀ i, α i) → Lex (Π₀ i, α i) not_mem_neLocus.mp (Finset.not_mem_of_lt_min hj <| by rwa [neLocus_comm]), hwit⟩ /-- The less-or-equal relation for the lexicographic ordering is decidable. -/ -irreducible_def Lex.decidableLE : @DecidableRel (Lex (Π₀ i, α i)) (· ≤ ·) := +irreducible_def Lex.decidableLE : DecidableRel (α := Lex (Π₀ i, α i)) (· ≤ ·) := lt_trichotomy_rec (fun h ↦ isTrue <| Or.inr h) (fun h ↦ isTrue <| Or.inl <| congr_arg _ h) fun h ↦ isFalse fun h' ↦ lt_irrefl _ (h.trans_le h') /-- The less-than relation for the lexicographic ordering is decidable. -/ -irreducible_def Lex.decidableLT : @DecidableRel (Lex (Π₀ i, α i)) (· < ·) := +irreducible_def Lex.decidableLT : DecidableRel (α := Lex (Π₀ i, α i)) (· < ·) := lt_trichotomy_rec (fun h ↦ isTrue h) (fun h ↦ isFalse h.not_lt) fun h ↦ isFalse h.asymm -- Porting note: Added `DecidableEq` for `LinearOrder`. diff --git a/Mathlib/Data/Finset/Defs.lean b/Mathlib/Data/Finset/Defs.lean index daad50dd4cc92..7acf539ddfefc 100644 --- a/Mathlib/Data/Finset/Defs.lean +++ b/Mathlib/Data/Finset/Defs.lean @@ -370,16 +370,16 @@ instance decidableDforallFinset {p : ∀ a ∈ s, Prop} [_hp : ∀ (a) (h : a -- Porting note: In lean3, `decidableDforallFinset` was picked up when decidability of `s ⊆ t` was -- needed. In lean4 it seems this is not the case. -instance instDecidableRelSubset [DecidableEq α] : @DecidableRel (Finset α) (· ⊆ ·) := +instance instDecidableRelSubset [DecidableEq α] : DecidableRel (α := Finset α) (· ⊆ ·) := fun _ _ ↦ decidableDforallFinset -instance instDecidableRelSSubset [DecidableEq α] : @DecidableRel (Finset α) (· ⊂ ·) := +instance instDecidableRelSSubset [DecidableEq α] : DecidableRel (α := Finset α) (· ⊂ ·) := fun _ _ ↦ instDecidableAnd -instance instDecidableLE [DecidableEq α] : @DecidableRel (Finset α) (· ≤ ·) := +instance instDecidableLE [DecidableEq α] : DecidableRel (α := Finset α) (· ≤ ·) := instDecidableRelSubset -instance instDecidableLT [DecidableEq α] : @DecidableRel (Finset α) (· < ·) := +instance instDecidableLT [DecidableEq α] : DecidableRel (α := Finset α) (· < ·) := instDecidableRelSSubset instance decidableDExistsFinset {p : ∀ a ∈ s, Prop} [_hp : ∀ (a) (h : a ∈ s), Decidable (p a h)] : diff --git a/Mathlib/Data/Finset/Sups.lean b/Mathlib/Data/Finset/Sups.lean index febbe67418d1c..04324a2636c5c 100644 --- a/Mathlib/Data/Finset/Sups.lean +++ b/Mathlib/Data/Finset/Sups.lean @@ -159,7 +159,7 @@ lemma sups_subset_self : s ⊻ s ⊆ s ↔ SupClosed (s : Set α) := sups_subset @[simp] lemma univ_sups_univ [Fintype α] : (univ : Finset α) ⊻ univ = univ := by simp -lemma filter_sups_le [@DecidableRel α (· ≤ ·)] (s t : Finset α) (a : α) : +lemma filter_sups_le [DecidableRel (α := α) (· ≤ ·)] (s t : Finset α) (a : α) : {b ∈ s ⊻ t | b ≤ a} = {b ∈ s | b ≤ a} ⊻ {b ∈ t | b ≤ a} := by simp only [← coe_inj, coe_filter, coe_sups, ← mem_coe, Set.sep_sups_le] @@ -304,7 +304,7 @@ lemma infs_self_subset : s ⊼ s ⊆ s ↔ InfClosed (s : Set α) := infs_subset @[simp] lemma univ_infs_univ [Fintype α] : (univ : Finset α) ⊼ univ = univ := by simp -lemma filter_infs_le [@DecidableRel α (· ≤ ·)] (s t : Finset α) (a : α) : +lemma filter_infs_le [DecidableRel (α := α) (· ≤ ·)] (s t : Finset α) (a : α) : {b ∈ s ⊼ t | a ≤ b} = {b ∈ s | a ≤ b} ⊼ {b ∈ t | a ≤ b} := by simp only [← coe_inj, coe_filter, coe_infs, ← mem_coe, Set.sep_infs_le] @@ -388,7 +388,8 @@ end Finset section DisjSups variable [DecidableEq α] -variable [SemilatticeSup α] [OrderBot α] [@DecidableRel α Disjoint] (s s₁ s₂ t t₁ t₂ u : Finset α) +variable [SemilatticeSup α] [OrderBot α] [DecidableRel (α := α) Disjoint] + (s s₁ s₂ t t₁ t₂ u : Finset α) /-- The finset of elements of the form `a ⊔ b` where `a ∈ s`, `b ∈ t` and `a` and `b` are disjoint. -/ @@ -486,7 +487,7 @@ open FinsetFamily section DistribLattice variable [DecidableEq α] -variable [DistribLattice α] [OrderBot α] [@DecidableRel α Disjoint] (s t u v : Finset α) +variable [DistribLattice α] [OrderBot α] [DecidableRel (α := α) Disjoint] (s t u v : Finset α) theorem disjSups_assoc : ∀ s t u : Finset α, s ○ t ○ u = s ○ (t ○ u) := by refine (associative_of_commutative_of_le inferInstance ?_).assoc diff --git a/Mathlib/Data/List/MinMax.lean b/Mathlib/Data/List/MinMax.lean index 2a7dbf99b60f2..5ed4618178ea5 100644 --- a/Mathlib/Data/List/MinMax.lean +++ b/Mathlib/Data/List/MinMax.lean @@ -89,7 +89,7 @@ end ArgAux section Preorder -variable [Preorder β] [@DecidableRel β (· < ·)] {f : α → β} {l : List α} {a m : α} +variable [Preorder β] [DecidableRel (α := β) (· < ·)] {f : α → β} {l : List α} {a m : α} /-- `argmax f l` returns `some a`, where `f a` is maximal among the elements of `l`, in the sense that there is no `b ∈ l` with `f a < f b`. If `a`, `b` are such that `f a = f b`, it returns @@ -243,7 +243,7 @@ section MaximumMinimum section Preorder -variable [Preorder α] [@DecidableRel α (· < ·)] {l : List α} {a m : α} +variable [Preorder α] [DecidableRel (α := α) (· < ·)] {l : List α} {a m : α} /-- `maximum l` returns a `WithBot α`, the largest element of `l` for nonempty lists, and `⊥` for `[]` -/ diff --git a/Mathlib/Data/Num/Basic.lean b/Mathlib/Data/Num/Basic.lean index 07f592d41ee94..54c9466d8c312 100644 --- a/Mathlib/Data/Num/Basic.lean +++ b/Mathlib/Data/Num/Basic.lean @@ -159,10 +159,10 @@ instance : LT PosNum := instance : LE PosNum := ⟨fun a b => ¬b < a⟩ -instance decidableLT : @DecidableRel PosNum (· < ·) +instance decidableLT : DecidableRel (α := PosNum) (· < ·) | a, b => by dsimp [LT.lt]; infer_instance -instance decidableLE : @DecidableRel PosNum (· ≤ ·) +instance decidableLE : DecidableRel (α := PosNum) (· ≤ ·) | a, b => by dsimp [LE.le]; infer_instance end PosNum @@ -270,10 +270,10 @@ instance : LT Num := instance : LE Num := ⟨fun a b => ¬b < a⟩ -instance decidableLT : @DecidableRel Num (· < ·) +instance decidableLT : DecidableRel (α := Num) (· < ·) | a, b => by dsimp [LT.lt]; infer_instance -instance decidableLE : @DecidableRel Num (· ≤ ·) +instance decidableLE : DecidableRel (α := Num) (· ≤ ·) | a, b => by dsimp [LE.le]; infer_instance /-- Converts a `Num` to a `ZNum`. -/ @@ -479,10 +479,10 @@ instance : LT ZNum := instance : LE ZNum := ⟨fun a b => ¬b < a⟩ -instance decidableLT : @DecidableRel ZNum (· < ·) +instance decidableLT : DecidableRel (α := ZNum) (· < ·) | a, b => by dsimp [LT.lt]; infer_instance -instance decidableLE : @DecidableRel ZNum (· ≤ ·) +instance decidableLE : DecidableRel (α := ZNum) (· ≤ ·) | a, b => by dsimp [LE.le]; infer_instance end ZNum diff --git a/Mathlib/Data/Ordmap/Ordnode.lean b/Mathlib/Data/Ordmap/Ordnode.lean index c4c775f222f1b..430a9571c3dbc 100644 --- a/Mathlib/Data/Ordmap/Ordnode.lean +++ b/Mathlib/Data/Ordmap/Ordnode.lean @@ -366,7 +366,7 @@ and should always be used instead of `Amem`. -/ def Amem [LE α] (x : α) : Ordnode α → Prop := Any fun y => x ≤ y ∧ y ≤ x -instance Amem.decidable [LE α] [@DecidableRel α (· ≤ ·)] (x : α) : +instance Amem.decidable [LE α] [DecidableRel (α := α) (· ≤ ·)] (x : α) : ∀ t, Decidable (Amem x t) := by dsimp [Amem]; infer_instance @@ -807,7 +807,7 @@ def ofAscList : List α → Ordnode α section -variable [LE α] [@DecidableRel α (· ≤ ·)] +variable [LE α] [DecidableRel (α := α) (· ≤ ·)] /-- O(log n). Does the set (approximately) contain the element `x`? That is, is there an element that is equivalent to `x` in the order? @@ -1220,7 +1220,7 @@ Equivalent elements are selected with a preference for smaller source elements. image (fun x ↦ x + 2) {1, 2, 4} = {3, 4, 6} image (fun x : ℕ ↦ x - 2) {1, 2, 4} = {0, 2} -/ -def image {α β} [LE β] [@DecidableRel β (· ≤ ·)] (f : α → β) (t : Ordnode α) : Ordnode β := +def image {α β} [LE β] [DecidableRel (α := β) (· ≤ ·)] (f : α → β) (t : Ordnode α) : Ordnode β := ofList (t.toList.map f) end diff --git a/Mathlib/Data/Ordmap/Ordset.lean b/Mathlib/Data/Ordmap/Ordset.lean index c2343c28c9c36..e1d9a4f8ef4b9 100644 --- a/Mathlib/Data/Ordmap/Ordset.lean +++ b/Mathlib/Data/Ordmap/Ordset.lean @@ -491,7 +491,7 @@ theorem equiv_iff {t₁ t₂ : Ordnode α} (h₁ : Sized t₁) (h₂ : Sized t /-! ### `mem` -/ -theorem pos_size_of_mem [LE α] [@DecidableRel α (· ≤ ·)] {x : α} {t : Ordnode α} (h : Sized t) +theorem pos_size_of_mem [LE α] [DecidableRel (α := α) (· ≤ ·)] {x : α} {t : Ordnode α} (h : Sized t) (h_mem : x ∈ t) : 0 < size t := by cases t; · { contradiction }; · { simp [h.1] } /-! ### `(find/erase/split)(Min/Max)` -/ @@ -564,7 +564,7 @@ theorem merge_node {ls ll lx lr rs rl rx rr} : /-! ### `insert` -/ -theorem dual_insert [Preorder α] [IsTotal α (· ≤ ·)] [@DecidableRel α (· ≤ ·)] (x : α) : +theorem dual_insert [Preorder α] [IsTotal α (· ≤ ·)] [DecidableRel (α := α) (· ≤ ·)] (x : α) : ∀ t : Ordnode α, dual (Ordnode.insert x t) = @Ordnode.insert αᵒᵈ _ _ x (dual t) | nil => rfl | node _ l y r => by @@ -1317,7 +1317,7 @@ theorem Valid.merge {l r} (hl : Valid l) (hr : Valid r) (sep : l.All fun x => r.All fun y => x < y) : Valid (@merge α l r) := (Valid'.merge_aux hl hr sep).1 -theorem insertWith.valid_aux [IsTotal α (· ≤ ·)] [@DecidableRel α (· ≤ ·)] (f : α → α) (x : α) +theorem insertWith.valid_aux [IsTotal α (· ≤ ·)] [DecidableRel (α := α) (· ≤ ·)] (f : α → α) (x : α) (hf : ∀ y, x ≤ y ∧ y ≤ x → x ≤ f y ∧ f y ≤ x) : ∀ {t o₁ o₂}, Valid' o₁ t o₂ → @@ -1346,28 +1346,28 @@ theorem insertWith.valid_aux [IsTotal α (· ≤ ·)] [@DecidableRel α (· ≤ exact (e.add_left _).add_right _ exact Or.inr ⟨_, e, h.3.1⟩ -theorem insertWith.valid [IsTotal α (· ≤ ·)] [@DecidableRel α (· ≤ ·)] (f : α → α) (x : α) +theorem insertWith.valid [IsTotal α (· ≤ ·)] [DecidableRel (α := α) (· ≤ ·)] (f : α → α) (x : α) (hf : ∀ y, x ≤ y ∧ y ≤ x → x ≤ f y ∧ f y ≤ x) {t} (h : Valid t) : Valid (insertWith f x t) := (insertWith.valid_aux _ _ hf h ⟨⟩ ⟨⟩).1 -theorem insert_eq_insertWith [@DecidableRel α (· ≤ ·)] (x : α) : +theorem insert_eq_insertWith [DecidableRel (α := α) (· ≤ ·)] (x : α) : ∀ t, Ordnode.insert x t = insertWith (fun _ => x) x t | nil => rfl | node _ l y r => by unfold Ordnode.insert insertWith; cases cmpLE x y <;> simp [insert_eq_insertWith] -theorem insert.valid [IsTotal α (· ≤ ·)] [@DecidableRel α (· ≤ ·)] (x : α) {t} (h : Valid t) : +theorem insert.valid [IsTotal α (· ≤ ·)] [DecidableRel (α := α) (· ≤ ·)] (x : α) {t} (h : Valid t) : Valid (Ordnode.insert x t) := by rw [insert_eq_insertWith]; exact insertWith.valid _ _ (fun _ _ => ⟨le_rfl, le_rfl⟩) h -theorem insert'_eq_insertWith [@DecidableRel α (· ≤ ·)] (x : α) : +theorem insert'_eq_insertWith [DecidableRel (α := α) (· ≤ ·)] (x : α) : ∀ t, insert' x t = insertWith id x t | nil => rfl | node _ l y r => by unfold insert' insertWith; cases cmpLE x y <;> simp [insert'_eq_insertWith] -theorem insert'.valid [IsTotal α (· ≤ ·)] [@DecidableRel α (· ≤ ·)] (x : α) {t} (h : Valid t) : - Valid (insert' x t) := by +theorem insert'.valid [IsTotal α (· ≤ ·)] [DecidableRel (α := α) (· ≤ ·)] + (x : α) {t} (h : Valid t) : Valid (insert' x t) := by rw [insert'_eq_insertWith]; exact insertWith.valid _ _ (fun _ => id) h theorem Valid'.map_aux {β} [Preorder β] {f : α → β} (f_strict_mono : StrictMono f) {t a₁ a₂} @@ -1404,7 +1404,7 @@ theorem map.valid {β} [Preorder β] {f : α → β} (f_strict_mono : StrictMono Valid (map f t) := (Valid'.map_aux f_strict_mono h).1 -theorem Valid'.erase_aux [@DecidableRel α (· ≤ ·)] (x : α) {t a₁ a₂} (h : Valid' a₁ t a₂) : +theorem Valid'.erase_aux [DecidableRel (α := α) (· ≤ ·)] (x : α) {t a₁ a₂} (h : Valid' a₁ t a₂) : Valid' a₁ (erase x t) a₂ ∧ Raised (erase x t).size t.size := by induction t generalizing a₁ a₂ with | nil => @@ -1438,10 +1438,10 @@ theorem Valid'.erase_aux [@DecidableRel α (· ≤ ·)] (x : α) {t a₁ a₂} ( exact t_r_size right; exists t_r.size; exact And.intro t_r_size h.bal.1 -theorem erase.valid [@DecidableRel α (· ≤ ·)] (x : α) {t} (h : Valid t) : Valid (erase x t) := +theorem erase.valid [DecidableRel (α := α) (· ≤ ·)] (x : α) {t} (h : Valid t) : Valid (erase x t) := (Valid'.erase_aux x h).1 -theorem size_erase_of_mem [@DecidableRel α (· ≤ ·)] {x : α} {t a₁ a₂} (h : Valid' a₁ t a₂) +theorem size_erase_of_mem [DecidableRel (α := α) (· ≤ ·)] {x : α} {t a₁ a₂} (h : Valid' a₁ t a₂) (h_mem : x ∈ t) : size (erase x t) = size t - 1 := by induction t generalizing a₁ a₂ with | nil => @@ -1528,22 +1528,22 @@ instance Empty.instDecidablePred : DecidablePred (@Empty α _) := /-- O(log n). Insert an element into the set, preserving balance and the BST property. If an equivalent element is already in the set, this replaces it. -/ -protected def insert [IsTotal α (· ≤ ·)] [@DecidableRel α (· ≤ ·)] (x : α) (s : Ordset α) : +protected def insert [IsTotal α (· ≤ ·)] [DecidableRel (α := α) (· ≤ ·)] (x : α) (s : Ordset α) : Ordset α := ⟨Ordnode.insert x s.1, insert.valid _ s.2⟩ -instance instInsert [IsTotal α (· ≤ ·)] [@DecidableRel α (· ≤ ·)] : Insert α (Ordset α) := +instance instInsert [IsTotal α (· ≤ ·)] [DecidableRel (α := α) (· ≤ ·)] : Insert α (Ordset α) := ⟨Ordset.insert⟩ /-- O(log n). Insert an element into the set, preserving balance and the BST property. If an equivalent element is already in the set, the set is returned as is. -/ -nonrec def insert' [IsTotal α (· ≤ ·)] [@DecidableRel α (· ≤ ·)] (x : α) (s : Ordset α) : +nonrec def insert' [IsTotal α (· ≤ ·)] [DecidableRel (α := α) (· ≤ ·)] (x : α) (s : Ordset α) : Ordset α := ⟨insert' x s.1, insert'.valid _ s.2⟩ section -variable [@DecidableRel α (· ≤ ·)] +variable [DecidableRel (α := α) (· ≤ ·)] /-- O(log n). Does the set contain the element `x`? That is, is there an element that is equivalent to `x` in the order? -/ @@ -1570,7 +1570,7 @@ end /-- O(log n). Remove an element from the set equivalent to `x`. Does nothing if there is no such element. -/ -def erase [@DecidableRel α (· ≤ ·)] (x : α) (s : Ordset α) : Ordset α := +def erase [DecidableRel (α := α) (· ≤ ·)] (x : α) (s : Ordset α) : Ordset α := ⟨Ordnode.erase x s.val, Ordnode.erase.valid x s.property⟩ /-- O(n). Map a function across a tree, without changing the structure. -/ diff --git a/Mathlib/Data/String/Basic.lean b/Mathlib/Data/String/Basic.lean index d77a4e02a1f33..b4a8b44ddfbf8 100644 --- a/Mathlib/Data/String/Basic.lean +++ b/Mathlib/Data/String/Basic.lean @@ -31,7 +31,7 @@ def ltb (s₁ s₂ : Iterator) : Bool := instance LT' : LT String := ⟨fun s₁ s₂ ↦ ltb s₁.iter s₂.iter⟩ -instance decidableLT : @DecidableRel String (· < ·) := by +instance decidableLT : DecidableRel (α := String) (· < ·) := by simp only [LT'] infer_instance -- short-circuit type class inference @@ -100,7 +100,7 @@ theorem lt_iff_toList_lt : ∀ {s₁ s₂ : String}, s₁ < s₂ ↔ s₁.toList instance LE : LE String := ⟨fun s₁ s₂ ↦ ¬s₂ < s₁⟩ -instance decidableLE : @DecidableRel String (· ≤ ·) := by +instance decidableLE : DecidableRel (α := String) (· ≤ ·) := by simp only [LE] infer_instance -- short-circuit type class inference diff --git a/Mathlib/Logic/Encodable/Basic.lean b/Mathlib/Logic/Encodable/Basic.lean index 60412ec201292..e1558f646d24a 100644 --- a/Mathlib/Logic/Encodable/Basic.lean +++ b/Mathlib/Logic/Encodable/Basic.lean @@ -592,7 +592,7 @@ section Quotient open Encodable Quotient -variable {α : Type*} {s : Setoid α} [@DecidableRel α (· ≈ ·)] [Encodable α] +variable {α : Type*} {s : Setoid α} [DecidableRel (α := α) (· ≈ ·)] [Encodable α] /-- Representative of an equivalence class. This is a computable version of `Quot.out` for a setoid on an encodable type. -/ diff --git a/Mathlib/NumberTheory/Zsqrtd/Basic.lean b/Mathlib/NumberTheory/Zsqrtd/Basic.lean index 914914db4f6cb..b7c9c0db4e4a0 100644 --- a/Mathlib/NumberTheory/Zsqrtd/Basic.lean +++ b/Mathlib/NumberTheory/Zsqrtd/Basic.lean @@ -558,7 +558,7 @@ instance decidableNonnegg (c d a b) : Decidable (Nonnegg c d a b) := by instance decidableNonneg : ∀ a : ℤ√d, Decidable (Nonneg a) | ⟨_, _⟩ => Zsqrtd.decidableNonnegg _ _ _ _ -instance decidableLE : @DecidableRel (ℤ√d) (· ≤ ·) := fun _ _ => decidableNonneg _ +instance decidableLE : DecidableRel (α := ℤ√d) (· ≤ ·) := fun _ _ => decidableNonneg _ open Int in theorem nonneg_cases : ∀ {a : ℤ√d}, Nonneg a → ∃ x y : ℕ, a = ⟨x, y⟩ ∨ a = ⟨x, -y⟩ ∨ a = ⟨-x, y⟩ diff --git a/Mathlib/Order/Antisymmetrization.lean b/Mathlib/Order/Antisymmetrization.lean index 75caf3593c355..2a5cced208b85 100644 --- a/Mathlib/Order/Antisymmetrization.lean +++ b/Mathlib/Order/Antisymmetrization.lean @@ -161,7 +161,7 @@ instance [WellFoundedLT α] : WellFoundedLT (Antisymmetrization α (· ≤ ·)) instance [WellFoundedGT α] : WellFoundedGT (Antisymmetrization α (· ≤ ·)) := wellFoundedGT_antisymmetrization_iff.mpr ‹_› -instance [@DecidableRel α (· ≤ ·)] [@DecidableRel α (· < ·)] [IsTotal α (· ≤ ·)] : +instance [DecidableRel (α := α) (· ≤ ·)] [DecidableRel (α := α) (· < ·)] [IsTotal α (· ≤ ·)] : LinearOrder (Antisymmetrization α (· ≤ ·)) := { instPartialOrderAntisymmetrization with le_total := fun a b => Quotient.inductionOn₂' a b <| total_of (· ≤ ·), diff --git a/Mathlib/Order/Basic.lean b/Mathlib/Order/Basic.lean index 8e5a610e2930f..d5004d18c24ab 100644 --- a/Mathlib/Order/Basic.lean +++ b/Mathlib/Order/Basic.lean @@ -289,7 +289,8 @@ section PartialOrder variable [PartialOrder α] {a b : α} -- See Note [decidable namespace] -protected theorem Decidable.le_iff_eq_or_lt [@DecidableRel α (· ≤ ·)] : a ≤ b ↔ a = b ∨ a < b := +protected theorem Decidable.le_iff_eq_or_lt [DecidableRel (α := α) (· ≤ ·)] : + a ≤ b ↔ a = b ∨ a < b := Decidable.le_iff_lt_or_eq.trans or_comm theorem le_iff_eq_or_lt : a ≤ b ↔ a = b ∨ a < b := le_iff_lt_or_eq.trans or_comm @@ -302,7 +303,7 @@ lemma eq_iff_not_lt_of_le (hab : a ≤ b) : a = b ↔ ¬ a < b := by simp [hab, alias LE.le.eq_iff_not_lt := eq_iff_not_lt_of_le -- See Note [decidable namespace] -protected theorem Decidable.eq_iff_le_not_lt [@DecidableRel α (· ≤ ·)] : +protected theorem Decidable.eq_iff_le_not_lt [DecidableRel (α := α) (· ≤ ·)] : a = b ↔ a ≤ b ∧ ¬a < b := ⟨fun h ↦ ⟨h.le, h ▸ lt_irrefl _⟩, fun ⟨h₁, h₂⟩ ↦ h₁.antisymm <| Decidable.byContradiction fun h₃ ↦ h₂ (h₁.lt_of_not_le h₃)⟩ @@ -1109,11 +1110,11 @@ instance preorder [Preorder α] (p : α → Prop) : Preorder (Subtype p) := instance partialOrder [PartialOrder α] (p : α → Prop) : PartialOrder (Subtype p) := PartialOrder.lift (fun (a : Subtype p) ↦ (a : α)) Subtype.coe_injective -instance decidableLE [Preorder α] [h : @DecidableRel α (· ≤ ·)] {p : α → Prop} : - @DecidableRel (Subtype p) (· ≤ ·) := fun a b ↦ h a b +instance decidableLE [Preorder α] [h : DecidableRel (α := α) (· ≤ ·)] {p : α → Prop} : + DecidableRel (α := Subtype p) (· ≤ ·) := fun a b ↦ h a b -instance decidableLT [Preorder α] [h : @DecidableRel α (· < ·)] {p : α → Prop} : - @DecidableRel (Subtype p) (· < ·) := fun a b ↦ h a b +instance decidableLT [Preorder α] [h : DecidableRel (α := α) (· < ·)] {p : α → Prop} : + DecidableRel (α := Subtype p) (· < ·) := fun a b ↦ h a b /-- A subtype of a linear order is a linear order. We explicitly give the proofs of decidable equality and decidable order in order to ensure the decidability instances are all definitionally diff --git a/Mathlib/Order/Chain.lean b/Mathlib/Order/Chain.lean index b15b044ceabf2..dc6fc7aa7d720 100644 --- a/Mathlib/Order/Chain.lean +++ b/Mathlib/Order/Chain.lean @@ -365,7 +365,7 @@ variable [PartialOrder α] theorem chain_lt (s : Flag α) : IsChain (· < ·) (s : Set α) := s.chain_le.lt_of_le -instance [@DecidableRel α (· ≤ ·)] [@DecidableRel α (· < ·)] (s : Flag α) : +instance [DecidableRel (α := α) (· ≤ ·)] [DecidableRel (α := α) (· < ·)] (s : Flag α) : LinearOrder s := { Subtype.partialOrder _ with le_total := fun a b => s.le_or_le a.2 b.2 diff --git a/Mathlib/Order/Compare.lean b/Mathlib/Order/Compare.lean index 5e29b6e29a069..763fb3ead16a0 100644 --- a/Mathlib/Order/Compare.lean +++ b/Mathlib/Order/Compare.lean @@ -25,16 +25,16 @@ variable {α β : Type*} /-- Like `cmp`, but uses a `≤` on the type instead of `<`. Given two elements `x` and `y`, returns a three-way comparison result `Ordering`. -/ -def cmpLE {α} [LE α] [@DecidableRel α (· ≤ ·)] (x y : α) : Ordering := +def cmpLE {α} [LE α] [DecidableRel (α := α) (· ≤ ·)] (x y : α) : Ordering := if x ≤ y then if y ≤ x then Ordering.eq else Ordering.lt else Ordering.gt -theorem cmpLE_swap {α} [LE α] [IsTotal α (· ≤ ·)] [@DecidableRel α (· ≤ ·)] (x y : α) : +theorem cmpLE_swap {α} [LE α] [IsTotal α (· ≤ ·)] [DecidableRel (α := α) (· ≤ ·)] (x y : α) : (cmpLE x y).swap = cmpLE y x := by by_cases xy : x ≤ y <;> by_cases yx : y ≤ x <;> simp [cmpLE, *, Ordering.swap] cases not_or_intro xy yx (total_of _ _ _) -theorem cmpLE_eq_cmp {α} [Preorder α] [IsTotal α (· ≤ ·)] [@DecidableRel α (· ≤ ·)] - [@DecidableRel α (· < ·)] (x y : α) : cmpLE x y = cmp x y := by +theorem cmpLE_eq_cmp {α} [Preorder α] [IsTotal α (· ≤ ·)] [DecidableRel (α := α) (· ≤ ·)] + [DecidableRel (α := α) (· < ·)] (x y : α) : cmpLE x y = cmp x y := by by_cases xy : x ≤ y <;> by_cases yx : y ≤ x <;> simp [cmpLE, lt_iff_le_not_le, *, cmp, cmpUsing] cases not_or_intro xy yx (total_of _ _ _) @@ -132,28 +132,29 @@ theorem Ordering.Compares.cmp_eq [LinearOrder α] {a b : α} {o : Ordering} (h : (cmp_compares a b).inj h @[simp] -theorem cmp_swap [Preorder α] [@DecidableRel α (· < ·)] (a b : α) : (cmp a b).swap = cmp b a := by +theorem cmp_swap [Preorder α] [DecidableRel (α := α) (· < ·)] (a b : α) : + (cmp a b).swap = cmp b a := by unfold cmp cmpUsing by_cases h : a < b <;> by_cases h₂ : b < a <;> simp [h, h₂, Ordering.swap] exact lt_asymm h h₂ @[simp] -theorem cmpLE_toDual [LE α] [@DecidableRel α (· ≤ ·)] (x y : α) : +theorem cmpLE_toDual [LE α] [DecidableRel (α := α) (· ≤ ·)] (x y : α) : cmpLE (toDual x) (toDual y) = cmpLE y x := rfl @[simp] -theorem cmpLE_ofDual [LE α] [@DecidableRel α (· ≤ ·)] (x y : αᵒᵈ) : +theorem cmpLE_ofDual [LE α] [DecidableRel (α := α) (· ≤ ·)] (x y : αᵒᵈ) : cmpLE (ofDual x) (ofDual y) = cmpLE y x := rfl @[simp] -theorem cmp_toDual [LT α] [@DecidableRel α (· < ·)] (x y : α) : +theorem cmp_toDual [LT α] [DecidableRel (α := α) (· < ·)] (x y : α) : cmp (toDual x) (toDual y) = cmp y x := rfl @[simp] -theorem cmp_ofDual [LT α] [@DecidableRel α (· < ·)] (x y : αᵒᵈ) : +theorem cmp_ofDual [LT α] [DecidableRel (α := α) (· < ·)] (x y : αᵒᵈ) : cmp (ofDual x) (ofDual y) = cmp y x := rfl diff --git a/Mathlib/Order/Defs/PartialOrder.lean b/Mathlib/Order/Defs/PartialOrder.lean index 377677f6d9855..d5bac42a0abd8 100644 --- a/Mathlib/Order/Defs/PartialOrder.lean +++ b/Mathlib/Order/Defs/PartialOrder.lean @@ -97,7 +97,7 @@ instance (priority := 900) : @Trans α α α GT.gt GE.ge GT.gt := ⟨gt_of_gt_of instance (priority := 900) : @Trans α α α GE.ge GT.gt GT.gt := ⟨gt_of_ge_of_gt⟩ /-- `<` is decidable if `≤` is. -/ -def decidableLTOfDecidableLE [@DecidableRel α (· ≤ ·)] : @DecidableRel α (· < ·) +def decidableLTOfDecidableLE [DecidableRel (α := α) (· ≤ ·)] : DecidableRel (α := α) (· < ·) | a, b => if hab : a ≤ b then if hba : b ≤ a then isFalse fun hab' => not_le_of_gt hab' hba @@ -129,7 +129,7 @@ lemma lt_of_le_of_ne : a ≤ b → a ≠ b → a < b := fun h₁ h₂ => lt_of_le_not_le h₁ <| mt (le_antisymm h₁) h₂ /-- Equality is decidable if `≤` is. -/ -def decidableEqOfDecidableLE [@DecidableRel α (· ≤ ·)] : DecidableEq α +def decidableEqOfDecidableLE [DecidableRel (α := α) (· ≤ ·)] : DecidableEq α | a, b => if hab : a ≤ b then if hba : b ≤ a then isTrue (le_antisymm hab hba) else isFalse fun heq => hba (heq ▸ le_refl _) @@ -137,7 +137,7 @@ def decidableEqOfDecidableLE [@DecidableRel α (· ≤ ·)] : DecidableEq α namespace Decidable -variable [@DecidableRel α (· ≤ ·)] +variable [DecidableRel (α := α) (· ≤ ·)] lemma lt_or_eq_of_le (hab : a ≤ b) : a < b ∨ a = b := if hba : b ≤ a then Or.inr (le_antisymm hab hba) else Or.inl (lt_of_le_not_le hab hba) diff --git a/Mathlib/Order/Interval/Basic.lean b/Mathlib/Order/Interval/Basic.lean index 55171625e65d7..291d014cb4e11 100644 --- a/Mathlib/Order/Interval/Basic.lean +++ b/Mathlib/Order/Interval/Basic.lean @@ -451,7 +451,7 @@ instance semilatticeSup : SemilatticeSup (Interval α) := section Decidable -variable [@DecidableRel α (· ≤ ·)] +variable [DecidableRel (α := α) (· ≤ ·)] instance lattice : Lattice (Interval α) := { Interval.semilatticeSup with @@ -574,7 +574,7 @@ section CompleteLattice variable [CompleteLattice α] -noncomputable instance completeLattice [@DecidableRel α (· ≤ ·)] : +noncomputable instance completeLattice [DecidableRel (α := α) (· ≤ ·)] : CompleteLattice (Interval α) := by classical exact @@ -659,7 +659,7 @@ noncomputable instance completeLattice [@DecidableRel α (· ≤ ·)] : } @[simp, norm_cast] -theorem coe_sInf [@DecidableRel α (· ≤ ·)] (S : Set (Interval α)) : +theorem coe_sInf [DecidableRel (α := α) (· ≤ ·)] (S : Set (Interval α)) : ↑(sInf S) = ⋂ s ∈ S, (s : Set α) := by classical -- Porting note: added -- Porting note: this `change` was @@ -678,12 +678,12 @@ theorem coe_sInf [@DecidableRel α (· ≤ ·)] (S : Set (Interval α)) : exact h fun s ha t hb => (hx _ ha).1.trans (hx _ hb).2 @[simp, norm_cast] -theorem coe_iInf [@DecidableRel α (· ≤ ·)] (f : ι → Interval α) : +theorem coe_iInf [DecidableRel (α := α) (· ≤ ·)] (f : ι → Interval α) : ↑(⨅ i, f i) = ⋂ i, (f i : Set α) := by simp [iInf] -- @[simp] -- Porting note: not in simpNF @[norm_cast] -theorem coe_iInf₂ [@DecidableRel α (· ≤ ·)] (f : ∀ i, κ i → Interval α) : +theorem coe_iInf₂ [DecidableRel (α := α) (· ≤ ·)] (f : ∀ i, κ i → Interval α) : ↑(⨅ (i) (j), f i j) = ⋂ (i) (j), (f i j : Set α) := by simp_rw [coe_iInf] end CompleteLattice diff --git a/Mathlib/Order/Interval/Finset/Box.lean b/Mathlib/Order/Interval/Finset/Box.lean index 9956d1f74bcb2..c84087a315416 100644 --- a/Mathlib/Order/Interval/Finset/Box.lean +++ b/Mathlib/Order/Interval/Finset/Box.lean @@ -62,7 +62,7 @@ open Finset namespace Prod variable {α β : Type*} [OrderedRing α] [OrderedRing β] [LocallyFiniteOrder α] [LocallyFiniteOrder β] - [DecidableEq α] [DecidableEq β] [@DecidableRel (α × β) (· ≤ ·)] + [DecidableEq α] [DecidableEq β] [DecidableRel (α := α × β) (· ≤ ·)] @[simp] lemma card_box_succ (n : ℕ) : #(box (n + 1) : Finset (α × β)) = diff --git a/Mathlib/Order/Interval/Finset/Defs.lean b/Mathlib/Order/Interval/Finset/Defs.lean index 730bf99b5e958..5ca55dc50e9eb 100644 --- a/Mathlib/Order/Interval/Finset/Defs.lean +++ b/Mathlib/Order/Interval/Finset/Defs.lean @@ -590,8 +590,8 @@ noncomputable def LocallyFiniteOrder.ofFiniteIcc (h : ∀ a b : α, (Set.Icc a b This is not an instance as it would not be defeq to better instances such as `Fin.locallyFiniteOrder`. -/ -abbrev Fintype.toLocallyFiniteOrder [Fintype α] [@DecidableRel α (· < ·)] - [@DecidableRel α (· ≤ ·)] : LocallyFiniteOrder α where +abbrev Fintype.toLocallyFiniteOrder [Fintype α] [DecidableRel (α := α) (· < ·)] + [DecidableRel (α := α) (· ≤ ·)] : LocallyFiniteOrder α where finsetIcc a b := (Set.Icc a b).toFinset finsetIco a b := (Set.Ico a b).toFinset finsetIoc a b := (Set.Ioc a b).toFinset @@ -790,7 +790,7 @@ end LocallyFiniteOrderTop /-! ### `Prod` -/ section LocallyFiniteOrder -variable [LocallyFiniteOrder α] [LocallyFiniteOrder β] [@DecidableRel (α × β) (· ≤ ·)] +variable [LocallyFiniteOrder α] [LocallyFiniteOrder β] [DecidableRel (α := α × β) (· ≤ ·)] instance Prod.instLocallyFiniteOrder : LocallyFiniteOrder (α × β) := LocallyFiniteOrder.ofIcc' (α × β) (fun x y ↦ Icc x.1 y.1 ×ˢ Icc x.2 y.2) fun a b x => by @@ -807,7 +807,7 @@ lemma Finset.card_Icc_prod (x y : α × β) : #(Icc x y) = #(Icc x.1 y.1) * #(Ic end LocallyFiniteOrder section LocallyFiniteOrderTop -variable [LocallyFiniteOrderTop α] [LocallyFiniteOrderTop β] [@DecidableRel (α × β) (· ≤ ·)] +variable [LocallyFiniteOrderTop α] [LocallyFiniteOrderTop β] [DecidableRel (α := α × β) (· ≤ ·)] instance Prod.instLocallyFiniteOrderTop : LocallyFiniteOrderTop (α × β) := LocallyFiniteOrderTop.ofIci' (α × β) (fun x => Ici x.1 ×ˢ Ici x.2) fun a x => by @@ -821,7 +821,7 @@ lemma Finset.card_Ici_prod (x : α × β) : #(Ici x) = #(Ici x.1) * #(Ici x.2) : end LocallyFiniteOrderTop section LocallyFiniteOrderBot -variable [LocallyFiniteOrderBot α] [LocallyFiniteOrderBot β] [@DecidableRel (α × β) (· ≤ ·)] +variable [LocallyFiniteOrderBot α] [LocallyFiniteOrderBot β] [DecidableRel (α := α × β) (· ≤ ·)] instance Prod.instLocallyFiniteOrderBot : LocallyFiniteOrderBot (α × β) := LocallyFiniteOrderBot.ofIic' (α × β) (fun x ↦ Iic x.1 ×ˢ Iic x.2) fun a x ↦ by @@ -836,7 +836,7 @@ end Preorder section Lattice variable [Lattice α] [Lattice β] [LocallyFiniteOrder α] [LocallyFiniteOrder β] - [@DecidableRel (α × β) (· ≤ ·)] + [DecidableRel (α := α × β) (· ≤ ·)] lemma Finset.uIcc_prod_def (x y : α × β) : uIcc x y = uIcc x.1 y.1 ×ˢ uIcc x.2 y.2 := rfl diff --git a/Mathlib/Order/SuccPred/Archimedean.lean b/Mathlib/Order/SuccPred/Archimedean.lean index d016540491343..0c89e2ea37b84 100644 --- a/Mathlib/Order/SuccPred/Archimedean.lean +++ b/Mathlib/Order/SuccPred/Archimedean.lean @@ -130,8 +130,8 @@ This isn't an instance due to a loop with `LinearOrder`. -/ -- See note [reducible non instances] abbrev IsSuccArchimedean.linearOrder [SuccOrder α] [IsSuccArchimedean α] - [DecidableEq α] [@DecidableRel α (· ≤ ·)] [@DecidableRel α (· < ·)] [IsDirected α (· ≥ ·)] : - LinearOrder α where + [DecidableEq α] [DecidableRel (α := α) (· ≤ ·)] [DecidableRel (α := α) (· < ·)] + [IsDirected α (· ≥ ·)] : LinearOrder α where le_total a b := have ⟨c, ha, hb⟩ := directed_of (· ≥ ·) a b le_total_of_codirected ha hb @@ -152,8 +152,8 @@ This isn't an instance due to a loop with `LinearOrder`. -/ -- See note [reducible non instances] abbrev IsPredArchimedean.linearOrder [PredOrder α] [IsPredArchimedean α] - [DecidableEq α] [@DecidableRel α (· ≤ ·)] [@DecidableRel α (· < ·)] [IsDirected α (· ≤ ·)] : - LinearOrder α := + [DecidableEq α] [DecidableRel (α := α) (· ≤ ·)] [DecidableRel (α := α) (· < ·)] + [IsDirected α (· ≤ ·)] : LinearOrder α := letI : LinearOrder αᵒᵈ := IsSuccArchimedean.linearOrder inferInstanceAs (LinearOrder αᵒᵈᵒᵈ) diff --git a/Mathlib/Order/WithBot.lean b/Mathlib/Order/WithBot.lean index da8552cf675af..7ceb5d577ecb0 100644 --- a/Mathlib/Order/WithBot.lean +++ b/Mathlib/Order/WithBot.lean @@ -445,13 +445,13 @@ instance distribLattice [DistribLattice α] : DistribLattice (WithBot α) := instance decidableEq [DecidableEq α] : DecidableEq (WithBot α) := inferInstanceAs <| DecidableEq (Option α) -instance decidableLE [LE α] [@DecidableRel α (· ≤ ·)] : @DecidableRel (WithBot α) (· ≤ ·) +instance decidableLE [LE α] [DecidableRel (α := α) (· ≤ ·)] : DecidableRel (α := WithBot α) (· ≤ ·) | none, _ => isTrue fun _ h => Option.noConfusion h | Option.some x, Option.some y => if h : x ≤ y then isTrue (coe_le_coe.2 h) else isFalse <| by simp [*] | Option.some x, none => isFalse fun h => by rcases h x rfl with ⟨y, ⟨_⟩, _⟩ -instance decidableLT [LT α] [@DecidableRel α (· < ·)] : @DecidableRel (WithBot α) (· < ·) +instance decidableLT [LT α] [DecidableRel (α := α) (· < ·)] : DecidableRel (α := WithBot α) (· < ·) | none, Option.some x => isTrue <| by exists x, rfl; rintro _ ⟨⟩ | Option.some x, Option.some y => if h : x < y then isTrue <| by simp [*] else isFalse <| by simp [*] @@ -1167,12 +1167,12 @@ instance distribLattice [DistribLattice α] : DistribLattice (WithTop α) := instance decidableEq [DecidableEq α] : DecidableEq (WithTop α) := inferInstanceAs <| DecidableEq (Option α) -instance decidableLE [LE α] [@DecidableRel α (· ≤ ·)] : - @DecidableRel (WithTop α) (· ≤ ·) := fun _ _ => +instance decidableLE [LE α] [DecidableRel (α := α) (· ≤ ·)] : + DecidableRel (α := WithTop α) (· ≤ ·) := fun _ _ => decidable_of_decidable_of_iff toDual_le_toDual_iff -instance decidableLT [LT α] [@DecidableRel α (· < ·)] : - @DecidableRel (WithTop α) (· < ·) := fun _ _ => +instance decidableLT [LT α] [DecidableRel (α := α) (· < ·)] : + DecidableRel (α := WithTop α) (· < ·) := fun _ _ => decidable_of_decidable_of_iff toDual_lt_toDual_iff instance isTotal_le [LE α] [IsTotal α (· ≤ ·)] : IsTotal (WithTop α) (· ≤ ·) := diff --git a/MathlibTest/instance_diamonds.lean b/MathlibTest/instance_diamonds.lean index 06ab59bc5c51d..48cf79849a3b7 100644 --- a/MathlibTest/instance_diamonds.lean +++ b/MathlibTest/instance_diamonds.lean @@ -226,8 +226,8 @@ end Polynomial section Subtype -- this diamond is the reason that `Fintype.toLocallyFiniteOrder` is not an instance -example {α} [Preorder α] [LocallyFiniteOrder α] [Fintype α] [@DecidableRel α (· < ·)] - [@DecidableRel α (· ≤ ·)] (p : α → Prop) [DecidablePred p] : +example {α} [Preorder α] [LocallyFiniteOrder α] [Fintype α] [DecidableRel (α := α) (· < ·)] + [DecidableRel (α := α) (· ≤ ·)] (p : α → Prop) [DecidablePred p] : Subtype.instLocallyFiniteOrder p = Fintype.toLocallyFiniteOrder := by fail_if_success rfl exact Subsingleton.elim _ _ From 013d296be2d78b736a8d271c0f5eabede7c04a7c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 11 Dec 2024 04:56:18 +0000 Subject: [PATCH 651/829] =?UTF-8?q?chore(SetTheory/Ordinal/Arithmetic):=20?= =?UTF-8?q?prove=20`n=20+=20=CF=89=20=3D=20=CF=89`=20earlier=20(#19803)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We use this to simplify the previous proof for `one_add_omega0`. --- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 73 +++++++++++------------ Mathlib/SetTheory/Ordinal/Principal.lean | 13 ---- 2 files changed, 36 insertions(+), 50 deletions(-) diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 4a4f0e4cbacc1..3f7e9f8d05724 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -612,27 +612,6 @@ theorem isLimit_sub {a b} (ha : IsLimit a) (h : b < a) : IsLimit (a - b) := by @[deprecated isLimit_sub (since := "2024-10-11")] alias sub_isLimit := isLimit_sub -theorem one_add_omega0 : 1 + ω = ω := by - refine le_antisymm ?_ (le_add_left _ _) - rw [omega0, ← lift_one.{0}, ← lift_add, lift_le, ← type_unit, ← type_sum_lex] - refine ⟨RelEmbedding.collapse (RelEmbedding.ofMonotone ?_ ?_)⟩ - · apply Sum.rec - · exact fun _ => 0 - · exact Nat.succ - · intro a b - 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 "No deprecation message was provided." (since := "2024-09-30")] -alias one_add_omega := one_add_omega0 - -@[simp] -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 "No deprecation message was provided." (since := "2024-09-30")] -alias one_add_of_omega_le := one_add_of_omega0_le - /-! ### Multiplication of ordinals -/ @@ -2329,24 +2308,8 @@ theorem lift_ofNat (n : ℕ) [n.AtLeastTwo] : lift.{u, v} (no_index (OfNat.ofNat n)) = OfNat.ofNat n := lift_natCast n -end Ordinal - /-! ### Properties of ω -/ - -namespace Cardinal - -open Ordinal - -@[simp] -theorem add_one_of_aleph0_le {c} (h : ℵ₀ ≤ c) : c + 1 = c := by - rw [add_comm, ← card_ord c, ← card_one, ← card_add, one_add_of_omega0_le] - rwa [← ord_aleph0, ord_le_ord] - -end Cardinal - -namespace Ordinal - theorem lt_add_of_limit {a b c : Ordinal.{u}} (h : IsLimit c) : a < b + c ↔ ∃ c' < c, a < b + c' := by -- Porting note: `bex_def` is required. @@ -2421,6 +2384,37 @@ theorem omega0_le_of_isLimit {o} (h : IsLimit o) : ω ≤ o := @[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias omega_le_of_isLimit := omega0_le_of_isLimit +theorem natCast_add_omega0 (n : ℕ) : n + ω = ω := by + refine le_antisymm (le_of_forall_lt fun a ha ↦ ?_) (le_add_left _ _) + obtain ⟨b, hb', hb⟩ := (lt_add_iff omega0_ne_zero).1 ha + obtain ⟨m, rfl⟩ := lt_omega0.1 hb' + apply hb.trans_lt + exact_mod_cast nat_lt_omega0 (n + m) + +theorem one_add_omega0 : 1 + ω = ω := + mod_cast natCast_add_omega0 1 + +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] +alias one_add_omega := one_add_omega0 + +theorem add_omega0 {a : Ordinal} (h : a < ω) : a + ω = ω := by + obtain ⟨n, rfl⟩ := lt_omega0.1 h + exact natCast_add_omega0 n + +@[deprecated (since := "2024-09-30")] +alias add_omega := add_omega0 + +@[simp] +theorem natCast_add_of_omega0_le {o} (h : ω ≤ o) (n : ℕ) : n + o = o := by + rw [← Ordinal.add_sub_cancel_of_le h, ← add_assoc, natCast_add_omega0] + +@[simp] +theorem one_add_of_omega0_le {o} (h : ω ≤ o) : 1 + o = o := + mod_cast natCast_add_of_omega0_le h 1 + +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] +alias one_add_of_omega_le := one_add_of_omega0_le + theorem isLimit_iff_omega0_dvd {a : Ordinal} : IsLimit a ↔ a ≠ 0 ∧ ω ∣ a := by refine ⟨fun l => ⟨l.ne_zero, ⟨a / ω, le_antisymm ?_ (mul_div_le _ _)⟩⟩, fun h => ?_⟩ · refine (limit_le l).2 fun x hx => le_of_lt ?_ @@ -2501,6 +2495,11 @@ namespace Cardinal open Ordinal +@[simp] +theorem add_one_of_aleph0_le {c} (h : ℵ₀ ≤ c) : c + 1 = c := by + rw [add_comm, ← card_ord c, ← card_one, ← card_add, one_add_of_omega0_le] + rwa [← ord_aleph0, ord_le_ord] + theorem isLimit_ord {c} (co : ℵ₀ ≤ c) : (ord c).IsLimit := by rw [isLimit_iff, isSuccPrelimit_iff_succ_lt] refine ⟨fun h => aleph0_ne_zero ?_, fun a => lt_imp_lt_of_le_imp_le fun h => ?_⟩ diff --git a/Mathlib/SetTheory/Ordinal/Principal.lean b/Mathlib/SetTheory/Ordinal/Principal.lean index 5ae52a926bb59..d993ba3d10462 100644 --- a/Mathlib/SetTheory/Ordinal/Principal.lean +++ b/Mathlib/SetTheory/Ordinal/Principal.lean @@ -195,19 +195,6 @@ theorem principal_add_iff_add_lt_ne_self : Principal (· + ·) a ↔ ∀ b < 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_omega0 (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_omega0_le (le_refl _)] - -@[deprecated (since := "2024-09-30")] -alias add_omega := add_omega0 - -@[simp] -theorem natCast_add_omega0 (n : ℕ) : n + ω = ω := - add_omega0 (nat_lt_omega0 n) - theorem principal_add_omega0 : Principal (· + ·) ω := principal_add_iff_add_left_eq_self.2 fun _ => add_omega0 From 9eb9253e89f61175c0284c6df0874b616e67e167 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Wed, 11 Dec 2024 06:10:06 +0000 Subject: [PATCH 652/829] =?UTF-8?q?feat(LinearAlgebra/Matrix):=20Prove=20`?= =?UTF-8?q?A=20*=20B=20=3D=201=20=E2=86=94=20B=20*=20A=20=3D=201`=20for=20?= =?UTF-8?q?square=20matrices=20over=20a=20commutative=20semiring=20(#19581?= =?UTF-8?q?)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR generalizes `Matrix.mul_eq_one_comm` from `CommRing` to `CommSemiring`. Co-authored-by: Eric Wieser --- Mathlib.lean | 1 + .../Algebra/BigOperators/Group/Finset.lean | 10 + Mathlib/Algebra/Module/End.lean | 13 +- Mathlib/Data/Matrix/ColumnRowPartitioned.lean | 3 +- Mathlib/Data/Matrix/Defs.lean | 6 + Mathlib/GroupTheory/Perm/Sign.lean | 26 ++ .../Matrix/Determinant/Basic.lean | 7 + .../Matrix/NonsingularInverse.lean | 50 +--- .../LinearAlgebra/Matrix/SemiringInverse.lean | 252 ++++++++++++++++++ 9 files changed, 316 insertions(+), 52 deletions(-) create mode 100644 Mathlib/LinearAlgebra/Matrix/SemiringInverse.lean diff --git a/Mathlib.lean b/Mathlib.lean index 2fdd9915d4878..81fbb6c85c38c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3421,6 +3421,7 @@ import Mathlib.LinearAlgebra.Matrix.PosDef import Mathlib.LinearAlgebra.Matrix.ProjectiveSpecialLinearGroup import Mathlib.LinearAlgebra.Matrix.Reindex import Mathlib.LinearAlgebra.Matrix.SchurComplement +import Mathlib.LinearAlgebra.Matrix.SemiringInverse import Mathlib.LinearAlgebra.Matrix.SesquilinearForm import Mathlib.LinearAlgebra.Matrix.SpecialLinearGroup import Mathlib.LinearAlgebra.Matrix.Spectrum diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index 88216bdc9801a..c31cf3896b7ed 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -2182,6 +2182,16 @@ theorem Units.coe_prod {M : Type*} [CommMonoid M] (f : α → Mˣ) (s : Finset (↑(∏ i ∈ s, f i) : M) = ∏ i ∈ s, (f i : M) := map_prod (Units.coeHom M) _ _ +@[to_additive (attr := simp)] +lemma IsUnit.prod_iff [CommMonoid β] : IsUnit (∏ a ∈ s, f a) ↔ ∀ a ∈ s, IsUnit (f a) := by + induction s using Finset.cons_induction with + | empty => simp + | cons a s ha hs => rw [Finset.prod_cons, IsUnit.mul_iff, hs, Finset.forall_mem_cons] + +@[to_additive] +lemma IsUnit.prod_univ_iff [Fintype α] [CommMonoid β] : IsUnit (∏ a, f a) ↔ ∀ a, IsUnit (f a) := by + simp + theorem nat_abs_sum_le {ι : Type*} (s : Finset ι) (f : ι → ℤ) : (∑ i ∈ s, f i).natAbs ≤ ∑ i ∈ s, (f i).natAbs := by induction s using Finset.cons_induction with diff --git a/Mathlib/Algebra/Module/End.lean b/Mathlib/Algebra/Module/End.lean index d11fb9a0448a3..426b42dfa3ea6 100644 --- a/Mathlib/Algebra/Module/End.lean +++ b/Mathlib/Algebra/Module/End.lean @@ -57,9 +57,20 @@ def smulAddHom : R →+ M →+ M := variable {R M} @[simp] -theorem smulAddHom_apply (r : R) (x : M) : smulAddHom R M r x = r • x := +theorem smulAddHom_apply : smulAddHom R M r x = r • x := rfl +variable {x} + +lemma IsAddUnit.smul_left [Monoid S] [DistribMulAction S M] (hx : IsAddUnit x) (s : S) : + IsAddUnit (s • x) := + hx.map (DistribMulAction.toAddMonoidHom M s) + +variable {r} (x) + +lemma IsAddUnit.smul_right (hr : IsAddUnit r) : IsAddUnit (r • x) := + hr.map (AddMonoidHom.flip (smulAddHom R M) x) + end AddCommMonoid section AddCommGroup diff --git a/Mathlib/Data/Matrix/ColumnRowPartitioned.lean b/Mathlib/Data/Matrix/ColumnRowPartitioned.lean index 67906c2b869f0..a8df024e8d340 100644 --- a/Mathlib/Data/Matrix/ColumnRowPartitioned.lean +++ b/Mathlib/Data/Matrix/ColumnRowPartitioned.lean @@ -4,9 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mohanad Ahmed -/ -import Mathlib.Data.Matrix.Basic import Mathlib.Data.Matrix.Block -import Mathlib.LinearAlgebra.Matrix.NonsingularInverse +import Mathlib.LinearAlgebra.Matrix.SemiringInverse /-! # Block Matrices from Rows and Columns diff --git a/Mathlib/Data/Matrix/Defs.lean b/Mathlib/Data/Matrix/Defs.lean index 55ae54ad81b76..7767c0b0b6d61 100644 --- a/Mathlib/Data/Matrix/Defs.lean +++ b/Mathlib/Data/Matrix/Defs.lean @@ -320,6 +320,12 @@ def ofAddEquiv [Add α] : (m → n → α) ≃+ Matrix m n α where @[simp] lemma coe_ofAddEquiv_symm [Add α] : ⇑(ofAddEquiv.symm : Matrix m n α ≃+ (m → n → α)) = of.symm := rfl +@[simp] lemma isAddUnit_iff [AddMonoid α] {A : Matrix m n α} : + IsAddUnit A ↔ ∀ i j, IsAddUnit (A i j) := by + simp_rw [isAddUnit_iff_exists, Classical.skolem, forall_and, + ← Matrix.ext_iff, add_apply, zero_apply] + rfl + end Matrix open Matrix diff --git a/Mathlib/GroupTheory/Perm/Sign.lean b/Mathlib/GroupTheory/Perm/Sign.lean index d8fe100a71871..9ef9e1f4486bd 100644 --- a/Mathlib/GroupTheory/Perm/Sign.lean +++ b/Mathlib/GroupTheory/Perm/Sign.lean @@ -9,6 +9,7 @@ import Mathlib.Algebra.Group.Submonoid.BigOperators import Mathlib.Algebra.Ring.Int.Units import Mathlib.Data.Finset.Fin import Mathlib.Data.Finset.Sort +import Mathlib.Data.Fintype.Perm import Mathlib.Data.Fintype.Prod import Mathlib.Data.Fintype.Sum import Mathlib.Data.Int.Order.Units @@ -597,4 +598,29 @@ end congr end SignType.sign +section Finset + +variable [Fintype α] + +/-- Permutations of a given sign. -/ +def ofSign (s : ℤˣ) : Finset (Perm α) := univ.filter (sign · = s) + +@[simp] +lemma mem_ofSign {s : ℤˣ} {σ : Perm α} : σ ∈ ofSign s ↔ σ.sign = s := by + rw [ofSign, mem_filter, and_iff_right (mem_univ σ)] + +lemma ofSign_disjoint : _root_.Disjoint (ofSign 1 : Finset (Perm α)) (ofSign (-1)) := by + rw [Finset.disjoint_iff_ne] + rintro σ hσ τ hτ rfl + rw [mem_ofSign] at hσ hτ + have := hσ.symm.trans hτ + contradiction + +lemma ofSign_disjUnion : + (ofSign 1).disjUnion (ofSign (-1)) ofSign_disjoint = (univ : Finset (Perm α)) := by + ext σ + simp_rw [mem_disjUnion, mem_ofSign, Int.units_eq_one_or, mem_univ] + +end Finset + end Equiv.Perm diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean index bd291498d65eb..1577138e86598 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean @@ -10,6 +10,7 @@ import Mathlib.Data.Matrix.RowCol import Mathlib.GroupTheory.GroupAction.Ring import Mathlib.GroupTheory.Perm.Fin import Mathlib.LinearAlgebra.Alternating.Basic +import Mathlib.LinearAlgebra.Matrix.SemiringInverse /-! # Determinant of a matrix @@ -62,6 +63,12 @@ theorem det_apply (M : Matrix n n R) : M.det = ∑ σ : Perm n, Equiv.Perm.sign theorem det_apply' (M : Matrix n n R) : M.det = ∑ σ : Perm n, ε σ * ∏ i, M (σ i) i := by simp [det_apply, Units.smul_def] +theorem det_eq_detp_sub_detp (M : Matrix n n R) : M.det = M.detp 1 - M.detp (-1) := by + rw [det_apply, ← Equiv.sum_comp (Equiv.inv (Perm n)), ← ofSign_disjUnion, sum_disjUnion] + simp_rw [inv_apply, sign_inv, sub_eq_add_neg, detp, ← sum_neg_distrib] + refine congr_arg₂ (· + ·) (sum_congr rfl fun σ hσ ↦ ?_) (sum_congr rfl fun σ hσ ↦ ?_) <;> + rw [mem_ofSign.mp hσ, ← Equiv.prod_comp σ] <;> simp + @[simp] theorem det_diagonal {d : n → R} : det (diagonal d) = ∏ i, d i := by rw [det_apply'] diff --git a/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean b/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean index 340f764a28faf..b279eda8cbc9c 100644 --- a/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean +++ b/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean @@ -6,6 +6,7 @@ Authors: Anne Baanen, Lu-Ming Zhang import Mathlib.Data.Matrix.Invertible import Mathlib.LinearAlgebra.FiniteDimensional.Defs import Mathlib.LinearAlgebra.Matrix.Adjugate +import Mathlib.LinearAlgebra.Matrix.SemiringInverse import Mathlib.LinearAlgebra.Matrix.Trace import Mathlib.LinearAlgebra.Matrix.ToLin @@ -108,29 +109,6 @@ def invertibleEquivDetInvertible : Invertible A ≃ Invertible A.det where left_inv _ := Subsingleton.elim _ _ right_inv _ := Subsingleton.elim _ _ -variable {A B} - -theorem mul_eq_one_comm : A * B = 1 ↔ B * A = 1 := - suffices ∀ A B : Matrix n n α, A * B = 1 → B * A = 1 from ⟨this A B, this B A⟩ - fun A B h => by - letI : Invertible B.det := detInvertibleOfLeftInverse _ _ h - letI : Invertible B := invertibleOfDetInvertible B - calc - B * A = B * A * (B * ⅟ B) := by rw [mul_invOf_self, Matrix.mul_one] - _ = B * (A * B * ⅟ B) := by simp only [Matrix.mul_assoc] - _ = B * ⅟ B := by rw [h, Matrix.one_mul] - _ = 1 := mul_invOf_self B - -variable (A B) - -/-- We can construct an instance of invertible A if A has a left inverse. -/ -def invertibleOfLeftInverse (h : B * A = 1) : Invertible A := - ⟨B, h, mul_eq_one_comm.mp h⟩ - -/-- We can construct an instance of invertible A if A has a right inverse. -/ -def invertibleOfRightInverse (h : A * B = 1) : Invertible A := - ⟨B, mul_eq_one_comm.mp h, h⟩ - /-- Given a proof that `A.det` has a constructive inverse, lift `A` to `(Matrix n n α)ˣ`-/ def unitOfDetInvertible [Invertible A.det] : (Matrix n n α)ˣ := @unitOfInvertible _ _ A (invertibleOfDetInvertible A) @@ -151,18 +129,6 @@ theorem isUnit_det_of_invertible [Invertible A] : IsUnit A.det := variable {A B} -theorem isUnit_of_left_inverse (h : B * A = 1) : IsUnit A := - ⟨⟨A, B, mul_eq_one_comm.mp h, h⟩, rfl⟩ - -theorem exists_left_inverse_iff_isUnit : (∃ B, B * A = 1) ↔ IsUnit A := - ⟨fun ⟨_, h⟩ ↦ isUnit_of_left_inverse h, fun h ↦ have := h.invertible; ⟨⅟A, invOf_mul_self' A⟩⟩ - -theorem isUnit_of_right_inverse (h : A * B = 1) : IsUnit A := - ⟨⟨A, B, h, mul_eq_one_comm.mp h⟩, rfl⟩ - -theorem exists_right_inverse_iff_isUnit : (∃ B, A * B = 1) ↔ IsUnit A := - ⟨fun ⟨_, h⟩ ↦ isUnit_of_right_inverse h, fun h ↦ have := h.invertible; ⟨⅟A, mul_invOf_self' A⟩⟩ - theorem isUnit_det_of_left_inverse (h : B * A = 1) : IsUnit A.det := @isUnit_of_invertible _ _ _ (detInvertibleOfLeftInverse _ _ h) @@ -177,20 +143,6 @@ theorem det_ne_zero_of_right_inverse [Nontrivial α] (h : A * B = 1) : A.det ≠ end Invertible - -section - -variable [Fintype m] [Fintype n] [DecidableEq m] [DecidableEq n] [CommRing α] - -/-- A version of `mul_eq_one_comm` that works for square matrices with rectangular types. -/ -theorem mul_eq_one_comm_of_equiv {A : Matrix m n α} {B : Matrix n m α} (e : m ≃ n) : - A * B = 1 ↔ B * A = 1 := by - refine (reindex e e).injective.eq_iff.symm.trans ?_ - rw [reindex_apply, reindex_apply, submatrix_one_equiv, ← submatrix_mul_equiv _ _ _ (.refl _), - mul_eq_one_comm, submatrix_mul_equiv, coe_refl, submatrix_id_id] - -end - section Inv variable [Fintype n] [DecidableEq n] [CommRing α] diff --git a/Mathlib/LinearAlgebra/Matrix/SemiringInverse.lean b/Mathlib/LinearAlgebra/Matrix/SemiringInverse.lean new file mode 100644 index 0000000000000..53f0096ec53d4 --- /dev/null +++ b/Mathlib/LinearAlgebra/Matrix/SemiringInverse.lean @@ -0,0 +1,252 @@ +/- +Copyright (c) 2024 Thomas Browning. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Thomas Browning +-/ +import Mathlib.Data.Fintype.Perm +import Mathlib.Data.Matrix.Mul +import Mathlib.GroupTheory.Perm.Sign + +/-! +# Nonsingular inverses over semirings + +This files proves `A * B = 1 ↔ B * A = 1` for square matrices over a commutative semiring. + +-/ + +open Equiv Equiv.Perm Finset + +variable {n m R : Type*} [Fintype m] [Fintype n] [DecidableEq m] [DecidableEq n] [CommSemiring R] + +variable (s : ℤˣ) (A B : Matrix n n R) (i j : n) + +namespace Matrix + +/-- The determinant, but only the terms of a given sign. -/ +def detp : R := ∑ σ ∈ ofSign s, ∏ k, A k (σ k) + +@[simp] +lemma detp_one_one : detp 1 (1 : Matrix n n R) = 1 := by + rw [detp, sum_eq_single_of_mem 1] + · simp [one_apply] + · simp [ofSign] + · rintro σ - hσ1 + obtain ⟨i, hi⟩ := not_forall.mp (mt Perm.ext_iff.mpr hσ1) + exact prod_eq_zero (mem_univ i) (one_apply_ne' hi) + +@[simp] +lemma detp_neg_one_one : detp (-1) (1 : Matrix n n R) = 0 := by + rw [detp, sum_eq_zero] + intro σ hσ + have hσ1 : σ ≠ 1 := by + contrapose! hσ + rw [hσ, mem_ofSign, sign_one] + decide + obtain ⟨i, hi⟩ := not_forall.mp (mt Perm.ext_iff.mpr hσ1) + exact prod_eq_zero (mem_univ i) (one_apply_ne' hi) + +/-- The adjugate matrix, but only the terms of a given sign. -/ +def adjp : Matrix n n R := + of fun i j ↦ ∑ σ ∈ (ofSign s).filter (· j = i), ∏ k ∈ {j}ᶜ, A k (σ k) + +lemma adjp_apply (i j : n) : + adjp s A i j = ∑ σ ∈ (ofSign s).filter (· j = i), ∏ k ∈ {j}ᶜ, A k (σ k) := + rfl + +theorem detp_mul : + detp 1 (A * B) + (detp 1 A * detp (-1) B + detp (-1) A * detp 1 B) = + detp (-1) (A * B) + (detp 1 A * detp 1 B + detp (-1) A * detp (-1) B) := by + have hf {s t} {σ : Perm n} (hσ : σ ∈ ofSign s) : + ofSign (t * s) = (ofSign t).map (mulRightEmbedding σ) := by + ext τ + simp_rw [mem_map, mulRightEmbedding_apply, ← eq_mul_inv_iff_mul_eq, exists_eq_right, + mem_ofSign, _root_.map_mul, _root_.map_inv, mul_inv_eq_iff_eq_mul, mem_ofSign.mp hσ] + have h {s t} : detp s A * detp t B = + ∑ σ ∈ ofSign s, ∑ τ ∈ ofSign (t * s), ∏ k, A k (σ k) * B (σ k) (τ k) := by + simp_rw [detp, sum_mul_sum, prod_mul_distrib] + refine sum_congr rfl fun σ hσ ↦ ?_ + simp_rw [hf hσ, sum_map, mulRightEmbedding_apply, Perm.mul_apply] + exact sum_congr rfl fun τ hτ ↦ (congr_arg (_ * ·) (Equiv.prod_comp σ _).symm) + let ι : Perm n ↪ (n → n) := ⟨_, coe_fn_injective⟩ + have hι {σ x} : ι σ x = σ x := rfl + let bij : Finset (n → n) := (disjUnion (ofSign 1) (ofSign (-1)) ofSign_disjoint).map ι + replace h (s) : detp s (A * B) = + ∑ σ ∈ bijᶜ, ∑ τ ∈ ofSign s, ∏ i : n, A i (σ i) * B (σ i) (τ i) + + (detp 1 A * detp s B + detp (-1) A * detp (-s) B) := by + simp_rw [h, neg_mul_neg, mul_one, detp, mul_apply, prod_univ_sum, Fintype.piFinset_univ] + rw [sum_comm, ← sum_compl_add_sum bij, sum_map, sum_disjUnion] + simp_rw [hι] + rw [h, h, neg_neg, add_assoc] + conv_rhs => rw [add_assoc] + refine congr_arg₂ (· + ·) (sum_congr rfl fun σ hσ ↦ ?_) (add_comm _ _) + replace hσ : ¬ Function.Injective σ := by + contrapose! hσ + rw [not_mem_compl, mem_map, ofSign_disjUnion] + exact ⟨Equiv.ofBijective σ hσ.bijective_of_finite, mem_univ _, rfl⟩ + obtain ⟨i, j, hσ, hij⟩ := Function.not_injective_iff.mp hσ + replace hσ k : σ (swap i j k) = σ k := by + rw [swap_apply_def] + split_ifs with h h <;> simp only [hσ, h] + rw [← mul_neg_one, hf (mem_ofSign.mpr (sign_swap hij)), sum_map] + simp_rw [prod_mul_distrib, mulRightEmbedding_apply, Perm.mul_apply] + refine sum_congr rfl fun τ hτ ↦ congr_arg (_ * ·) ?_ + rw [← Equiv.prod_comp (swap i j)] + simp only [hσ] + +theorem mul_adjp_apply_eq : (A * adjp s A) i i = detp s A := by + have key := sum_fiberwise_eq_sum_filter (ofSign s) univ (· i) fun σ ↦ ∏ k, A k (σ k) + simp_rw [mem_univ, filter_True] at key + simp_rw [mul_apply, adjp_apply, mul_sum, detp, ← key] + refine sum_congr rfl fun x hx ↦ sum_congr rfl fun σ hσ ↦ ?_ + rw [← prod_mul_prod_compl ({i} : Finset n), prod_singleton, (mem_filter.mp hσ).2] + +theorem mul_adjp_apply_ne (h : i ≠ j) : (A * adjp 1 A) i j = (A * adjp (-1) A) i j := by + simp_rw [mul_apply, adjp_apply, mul_sum, sum_sigma'] + let f : (Σ x : n, Perm n) → (Σ x : n, Perm n) := fun ⟨x, σ⟩ ↦ ⟨σ i, σ * swap i j⟩ + let t s : Finset (Σ x : n, Perm n) := univ.sigma fun x ↦ (ofSign s).filter fun σ ↦ σ j = x + have hf {s} : ∀ p ∈ t s, f (f p) = p := by + intro ⟨x, σ⟩ hp + rw [mem_sigma, mem_filter, mem_ofSign] at hp + simp_rw [f, Perm.mul_apply, swap_apply_left, hp.2.2, mul_swap_mul_self] + refine sum_bij' (fun p _ ↦ f p) (fun p _ ↦ f p) ?_ ?_ hf hf ?_ + · intro ⟨x, σ⟩ hp + rw [mem_sigma, mem_filter, mem_ofSign] at hp ⊢ + rw [Perm.mul_apply, sign_mul, hp.2.1, sign_swap h, swap_apply_right] + exact ⟨mem_univ (σ i), rfl, rfl⟩ + · intro ⟨x, σ⟩ hp + rw [mem_sigma, mem_filter, mem_ofSign] at hp ⊢ + rw [Perm.mul_apply, sign_mul, hp.2.1, sign_swap h, swap_apply_right] + exact ⟨mem_univ (σ i), rfl, rfl⟩ + · intro ⟨x, σ⟩ hp + rw [mem_sigma, mem_filter, mem_ofSign] at hp + have key : ({j}ᶜ : Finset n) = disjUnion ({i} : Finset n) ({i, j} : Finset n)ᶜ (by simp) := by + rw [singleton_disjUnion, cons_eq_insert, compl_insert, insert_erase] + rwa [mem_compl, mem_singleton] + simp_rw [key, prod_disjUnion, prod_singleton, Perm.mul_apply, swap_apply_left, ← mul_assoc] + rw [mul_comm (A i x) (A i (σ i)), hp.2.2] + refine congr_arg _ (prod_congr rfl fun x hx ↦ ?_) + rw [mem_compl, mem_insert, mem_singleton, not_or] at hx + rw [swap_apply_of_ne_of_ne hx.1 hx.2] + +theorem mul_adjp_add_detp : A * adjp 1 A + detp (-1) A • 1 = A * adjp (-1) A + detp 1 A • 1 := by + ext i j + rcases eq_or_ne i j with rfl | h <;> simp_rw [add_apply, smul_apply, smul_eq_mul] + · simp_rw [mul_adjp_apply_eq, one_apply_eq, mul_one, add_comm] + · simp_rw [mul_adjp_apply_ne A i j h, one_apply_ne h, mul_zero] + +variable {A B} + +theorem isAddUnit_mul (hAB : A * B = 1) (i j k : n) (hij : i ≠ j) : IsAddUnit (A i k * B k j) := by + revert k + rw [← IsAddUnit.sum_univ_iff, ← mul_apply, hAB, one_apply_ne hij] + exact isAddUnit_zero + +theorem isAddUnit_detp_mul_detp (hAB : A * B = 1) : + IsAddUnit (detp 1 A * detp (-1) B + detp (-1) A * detp 1 B) := by + suffices h : ∀ {s t}, s ≠ t → IsAddUnit (detp s A * detp t B) from + (h (by decide)).add (h (by decide)) + intro s t h + simp_rw [detp, sum_mul_sum, IsAddUnit.sum_iff] + intro σ hσ τ hτ + rw [mem_ofSign] at hσ hτ + rw [← hσ, ← hτ, ← sign_inv] at h + replace h := ne_of_apply_ne sign h + rw [ne_eq, eq_comm, eq_inv_iff_mul_eq_one, eq_comm] at h + simp_rw [Equiv.ext_iff, not_forall, Perm.mul_apply, Perm.one_apply] at h + obtain ⟨k, hk⟩ := h + rw [mul_comm, ← Equiv.prod_comp σ, mul_comm, ← prod_mul_distrib, + ← mul_prod_erase univ _ (mem_univ k), ← smul_eq_mul] + exact (isAddUnit_mul hAB k (τ (σ k)) (σ k) hk).smul_right _ + +theorem isAddUnit_detp_smul_mul_adjp (hAB : A * B = 1) : + IsAddUnit (detp 1 A • (B * adjp (-1) B) + detp (-1) A • (B * adjp 1 B)) := by + suffices h : ∀ {s t}, s ≠ t → IsAddUnit (detp s A • (B * adjp t B)) from + (h (by decide)).add (h (by decide)) + intro s t h + rw [isAddUnit_iff] + intro i j + simp_rw [smul_apply, smul_eq_mul, mul_apply, detp, adjp_apply, mul_sum, sum_mul, + IsAddUnit.sum_iff] + intro k hk σ hσ τ hτ + rw [mem_filter] at hσ + rw [mem_ofSign] at hσ hτ + rw [← hσ.1, ← hτ, ← sign_inv] at h + replace h := ne_of_apply_ne sign h + rw [ne_eq, eq_comm, eq_inv_iff_mul_eq_one] at h + obtain ⟨l, hl1, hl2⟩ := exists_ne_of_one_lt_card (one_lt_card_support_of_ne_one h) (τ⁻¹ j) + rw [mem_support, ne_comm] at hl1 + rw [ne_eq, ← mem_singleton, ← mem_compl] at hl2 + rw [← prod_mul_prod_compl {τ⁻¹ j}, mul_mul_mul_comm, mul_comm, ← smul_eq_mul] + apply IsAddUnit.smul_right + have h0 : ∀ k, k ∈ ({τ⁻¹ j} : Finset n)ᶜ ↔ τ k ∈ ({j} : Finset n)ᶜ := by + simp [show τ⁻¹ = τ.symm from rfl, eq_symm_apply] + rw [← prod_equiv τ h0 fun _ _ ↦ rfl, ← prod_mul_distrib, ← mul_prod_erase _ _ hl2, ← smul_eq_mul] + exact (isAddUnit_mul hAB l (σ (τ l)) (τ l) hl1).smul_right _ + +theorem detp_smul_add_adjp (hAB : A * B = 1) : + detp 1 B • A + adjp (-1) B = detp (-1) B • A + adjp 1 B := by + have key := congr(A * $(mul_adjp_add_detp B)) + simp_rw [mul_add, ← mul_assoc, hAB, one_mul, mul_smul, mul_one] at key + rwa [add_comm, eq_comm, add_comm] + +theorem detp_smul_adjp (hAB : A * B = 1) : + A + (detp 1 A • adjp (-1) B + detp (-1) A • adjp 1 B) = + detp 1 A • adjp 1 B + detp (-1) A • adjp (-1) B := by + have h0 := detp_mul A B + rw [hAB, detp_one_one, detp_neg_one_one, zero_add] at h0 + have h := detp_smul_add_adjp hAB + replace h := congr(detp 1 A • $h + detp (-1) A • $h.symm) + simp only [smul_add, smul_smul] at h + rwa [add_add_add_comm, ← add_smul, add_add_add_comm, ← add_smul, ← h0, add_smul, one_smul, + add_comm A, add_assoc, ((isAddUnit_detp_mul_detp hAB).smul_right _).add_right_inj] at h + +theorem mul_eq_one_comm : A * B = 1 ↔ B * A = 1 := by + suffices h : ∀ A B : Matrix n n R, A * B = 1 → B * A = 1 from ⟨h A B, h B A⟩ + intro A B hAB + have h0 := detp_mul A B + rw [hAB, detp_one_one, detp_neg_one_one, zero_add] at h0 + replace h := congr(B * $(detp_smul_adjp hAB)) + simp only [mul_add, mul_smul, add_assoc] at h + replace h := congr($h + (detp 1 A * detp (-1) B + detp (-1) A * detp 1 B) • 1) + simp_rw [add_smul, ← smul_smul] at h + rwa [add_assoc, add_add_add_comm, ← smul_add, ← smul_add, + add_add_add_comm, ← smul_add, ← smul_add, smul_add, smul_add, + mul_adjp_add_detp, smul_add, ← mul_adjp_add_detp, smul_add, ← smul_add, ← smul_add, + add_add_add_comm, smul_smul, smul_smul, ← add_smul, ← h0, + add_smul, one_smul, ← add_assoc _ 1, add_comm _ 1, add_assoc, + smul_add, smul_add, add_add_add_comm, smul_smul, smul_smul, ← add_smul, + ((isAddUnit_detp_smul_mul_adjp hAB).add + ((isAddUnit_detp_mul_detp hAB).smul_right _)).add_left_inj] at h + +variable (A B) + +/-- We can construct an instance of invertible A if A has a left inverse. -/ +def invertibleOfLeftInverse (h : B * A = 1) : Invertible A := + ⟨B, h, mul_eq_one_comm.mp h⟩ + +/-- We can construct an instance of invertible A if A has a right inverse. -/ +def invertibleOfRightInverse (h : A * B = 1) : Invertible A := + ⟨B, mul_eq_one_comm.mp h, h⟩ + +variable {A B} + +theorem isUnit_of_left_inverse (h : B * A = 1) : IsUnit A := + ⟨⟨A, B, mul_eq_one_comm.mp h, h⟩, rfl⟩ + +theorem exists_left_inverse_iff_isUnit : (∃ B, B * A = 1) ↔ IsUnit A := + ⟨fun ⟨_, h⟩ ↦ isUnit_of_left_inverse h, fun h ↦ have := h.invertible; ⟨⅟A, invOf_mul_self' A⟩⟩ + +theorem isUnit_of_right_inverse (h : A * B = 1) : IsUnit A := + ⟨⟨A, B, h, mul_eq_one_comm.mp h⟩, rfl⟩ + +theorem exists_right_inverse_iff_isUnit : (∃ B, A * B = 1) ↔ IsUnit A := + ⟨fun ⟨_, h⟩ ↦ isUnit_of_right_inverse h, fun h ↦ have := h.invertible; ⟨⅟A, mul_invOf_self' A⟩⟩ + +/-- A version of `mul_eq_one_comm` that works for square matrices with rectangular types. -/ +theorem mul_eq_one_comm_of_equiv {A : Matrix m n R} {B : Matrix n m R} (e : m ≃ n) : + A * B = 1 ↔ B * A = 1 := by + refine (reindex e e).injective.eq_iff.symm.trans ?_ + rw [reindex_apply, reindex_apply, submatrix_one_equiv, ← submatrix_mul_equiv _ _ _ (.refl _), + mul_eq_one_comm, submatrix_mul_equiv, coe_refl, submatrix_id_id] + +end Matrix From 097f9d2dc4052c4d70552604b5b36102e3ea710f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 11 Dec 2024 07:05:04 +0000 Subject: [PATCH 653/829] chore(SetTheory/Ordinal/Arithmetic): move `add_mul_limit` (#19805) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This has nothing to do with `ω` and can be moved much earlier in the file. We also private the auxiliary lemma `add_mul_limit_aux`. --- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 50 +++++++++++------------ 1 file changed, 24 insertions(+), 26 deletions(-) diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 3f7e9f8d05724..bd4d0884f8560 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -825,6 +825,30 @@ theorem smul_eq_mul : ∀ (n : ℕ) (a : Ordinal), n • a = a * n | 0, a => by rw [zero_nsmul, Nat.cast_zero, mul_zero] | n + 1, a => by rw [succ_nsmul, Nat.cast_add, mul_add, Nat.cast_one, mul_one, smul_eq_mul n] +private 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 + ((mul_le_of_limit l).2 fun c' h => by + apply (mul_le_mul_left' (le_succ c') _).trans + rw [IH _ h] + apply (add_le_add_left _ _).trans + · rw [← mul_succ] + exact mul_le_mul_left' (succ_le_of_lt <| l.succ_lt h) _ + · rw [← ba] + exact le_add_right _ _) + (mul_le_mul_right' (le_add_right _ _) _) + +theorem add_mul_succ {a b : Ordinal} (c) (ba : b + a = a) : (a + b) * succ c = a * succ c + b := by + induction c using limitRecOn with + | H₁ => simp only [succ_zero, mul_one] + | H₂ c IH => + rw [mul_succ, IH, ← add_assoc, add_assoc _ b, ba, ← mul_succ] + | H₃ c l IH => + rw [mul_succ, add_mul_limit_aux ba l IH, mul_succ, add_assoc] + +theorem add_mul_limit {a b c : Ordinal} (ba : b + a = a) (l : IsLimit c) : (a + b) * c = a * c := + add_mul_limit_aux ba l fun c' _ => add_mul_succ c' ba + /-! ### Division on ordinals -/ @@ -2433,32 +2457,6 @@ theorem isLimit_iff_omega0_dvd {a : Ordinal} : IsLimit a ↔ a ≠ 0 ∧ ω ∣ @[deprecated "No deprecation message was provided." (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 - ((mul_le_of_limit l).2 fun c' h => by - apply (mul_le_mul_left' (le_succ c') _).trans - rw [IH _ h] - apply (add_le_add_left _ _).trans - · rw [← mul_succ] - exact mul_le_mul_left' (succ_le_of_lt <| l.succ_lt h) _ - · rw [← ba] - exact le_add_right _ _) - (mul_le_mul_right' (le_add_right _ _) _) - -theorem add_mul_succ {a b : Ordinal} (c) (ba : b + a = a) : (a + b) * succ c = a * succ c + b := by - induction c using limitRecOn with - | H₁ => simp only [succ_zero, mul_one] - | H₂ c IH => - rw [mul_succ, IH, ← add_assoc, add_assoc _ b, ba, ← mul_succ] - | H₃ c l IH => - -- Porting note: Unused. - -- have := add_mul_limit_aux ba l IH - rw [mul_succ, add_mul_limit_aux ba l IH, mul_succ, add_assoc] - -theorem add_mul_limit {a b c : Ordinal} (ba : b + a = a) (l : IsLimit c) : (a + b) * c = a * c := - add_mul_limit_aux ba l fun c' _ => add_mul_succ c' ba - theorem IsNormal.apply_omega0 {f : Ordinal.{u} → Ordinal.{v}} (hf : IsNormal f) : ⨆ n : ℕ, f n = f ω := by rw [← iSup_natCast, hf.map_iSup] From 5f72b2e28cefb818887e4a7b7de72b6062fa1fc1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Martin=20Dvo=C5=99=C3=A1k?= Date: Wed, 11 Dec 2024 08:06:25 +0000 Subject: [PATCH 654/829] docs(LinearAlgebra/InvariantBasisNumber): first letter of a sentence (#19861) --- Mathlib/LinearAlgebra/InvariantBasisNumber.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/LinearAlgebra/InvariantBasisNumber.lean b/Mathlib/LinearAlgebra/InvariantBasisNumber.lean index 2979de0ca709c..fd7f09efc5db7 100644 --- a/Mathlib/LinearAlgebra/InvariantBasisNumber.lean +++ b/Mathlib/LinearAlgebra/InvariantBasisNumber.lean @@ -24,10 +24,10 @@ Let `R` be a (not necessary commutative) ring. It is also useful to consider the following stronger conditions: -- the *rank condition*, witnessed by the type class `RankCondition R`, states that - the existence of a surjective linear map `(Fin n → R) →ₗ[R] (Fin m → R)` implies `m ≤ n` +- The *rank condition*, witnessed by the type class `RankCondition R`, states that + the existence of a surjective linear map `(Fin n → R) →ₗ[R] (Fin m → R)` implies `m ≤ n`. -- the *strong rank condition*, witnessed by the type class `StrongRankCondition R`, states +- The *strong rank condition*, witnessed by the type class `StrongRankCondition R`, states that the existence of an injective linear map `(Fin n → R) →ₗ[R] (Fin m → R)` implies `n ≤ m`. From afdb9499da5095474266bcb3b723120eaf923aec Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Martin=20Dvo=C5=99=C3=A1k?= Date: Wed, 11 Dec 2024 08:26:29 +0000 Subject: [PATCH 655/829] chore: renaming all `Column` to `Col` (#19825) * `Matrix.fromColumns` becomes `Matrix.fromCols` * `Matrix.toColumns` becomes `Matrix.toCols` * `Matrix.updateColumn` becomes `Matrix.updateCol` Moves: - `*fromColumns*` -> `*fromCols*` - `*toColumns*` -> `*toCols*` - `*updateColumn*` -> `*updateCol*` Discussion: https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/Matrix.2EfromRows.20and.20Matrix.2EfromColumns --- Mathlib/Data/Matrix/ColumnRowPartitioned.lean | 175 ++++++++++++------ Mathlib/Data/Matrix/Notation.lean | 9 +- Mathlib/Data/Matrix/RowCol.lean | 123 +++++++----- Mathlib/LinearAlgebra/Determinant.lean | 4 +- Mathlib/LinearAlgebra/Matrix/Adjugate.lean | 40 ++-- Mathlib/LinearAlgebra/Matrix/Basis.lean | 4 +- .../Matrix/Determinant/Basic.lean | 44 +++-- .../Matrix/Determinant/TotallyUnimodular.lean | 31 ++-- Mathlib/LinearAlgebra/Matrix/Permanent.lean | 15 +- .../RingTheory/DedekindDomain/Different.lean | 2 +- Mathlib/RingTheory/Discriminant.lean | 4 +- Mathlib/Topology/Instances/Matrix.lean | 11 +- scripts/nolints_prime_decls.txt | 2 +- 13 files changed, 294 insertions(+), 170 deletions(-) diff --git a/Mathlib/Data/Matrix/ColumnRowPartitioned.lean b/Mathlib/Data/Matrix/ColumnRowPartitioned.lean index a8df024e8d340..c583ebb1d3b5e 100644 --- a/Mathlib/Data/Matrix/ColumnRowPartitioned.lean +++ b/Mathlib/Data/Matrix/ColumnRowPartitioned.lean @@ -11,7 +11,7 @@ import Mathlib.LinearAlgebra.Matrix.SemiringInverse This file provides the basic definitions of matrices composed from columns and rows. The concatenation of two matrices with the same row indices can be expressed as -`A = fromColumns A₁ A₂` the concatenation of two matrices with the same column indices +`A = fromCols A₁ A₂` the concatenation of two matrices with the same column indices can be expressed as `B = fromRows B₁ B₂`. We then provide a few lemmas that deal with the products of these with each other and @@ -33,14 +33,14 @@ def fromRows (A₁ : Matrix m₁ n R) (A₂ : Matrix m₂ n R) : Matrix (m₁ /-- Concatenate together two matrices B₁[m × n₁] and B₂[m × n₂] with the same rows (M) to get a bigger matrix indexed by [m × (n₁ ⊕ n₂)] -/ -def fromColumns (B₁ : Matrix m n₁ R) (B₂ : Matrix m n₂ R) : Matrix m (n₁ ⊕ n₂) R := +def fromCols (B₁ : Matrix m n₁ R) (B₂ : Matrix m n₂ R) : Matrix m (n₁ ⊕ n₂) R := of fun i => Sum.elim (B₁ i) (B₂ i) /-- Given a column partitioned matrix extract the first column -/ -def toColumns₁ (A : Matrix m (n₁ ⊕ n₂) R) : Matrix m n₁ R := of fun i j => (A i (Sum.inl j)) +def toCols₁ (A : Matrix m (n₁ ⊕ n₂) R) : Matrix m n₁ R := of fun i j => (A i (Sum.inl j)) /-- Given a column partitioned matrix extract the second column -/ -def toColumns₂ (A : Matrix m (n₁ ⊕ n₂) R) : Matrix m n₂ R := of fun i j => (A i (Sum.inr j)) +def toCols₂ (A : Matrix m (n₁ ⊕ n₂) R) : Matrix m n₂ R := of fun i j => (A i (Sum.inr j)) /-- Given a row partitioned matrix extract the first row -/ def toRows₁ (A : Matrix (m₁ ⊕ m₂) n R) : Matrix m₁ n R := of fun i j => (A (Sum.inl i) j) @@ -48,6 +48,10 @@ def toRows₁ (A : Matrix (m₁ ⊕ m₂) n R) : Matrix m₁ n R := of fun i j = /-- Given a row partitioned matrix extract the second row -/ def toRows₂ (A : Matrix (m₁ ⊕ m₂) n R) : Matrix m₂ n R := of fun i j => (A (Sum.inr i) j) +@[deprecated (since := "2024-12-11")] alias fromColumns := fromCols +@[deprecated (since := "2024-12-11")] alias toColumns₁ := toCols₁ +@[deprecated (since := "2024-12-11")] alias toColumns₂ := toCols₂ + @[simp] lemma fromRows_apply_inl (A₁ : Matrix m₁ n R) (A₂ : Matrix m₂ n R) (i : m₁) (j : n) : (fromRows A₁ A₂) (Sum.inl i) j = A₁ i j := rfl @@ -57,12 +61,16 @@ lemma fromRows_apply_inr (A₁ : Matrix m₁ n R) (A₂ : Matrix m₂ n R) (i : (fromRows A₁ A₂) (Sum.inr i) j = A₂ i j := rfl @[simp] -lemma fromColumns_apply_inl (A₁ : Matrix m n₁ R) (A₂ : Matrix m n₂ R) (i : m) (j : n₁) : - (fromColumns A₁ A₂) i (Sum.inl j) = A₁ i j := rfl +lemma fromCols_apply_inl (A₁ : Matrix m n₁ R) (A₂ : Matrix m n₂ R) (i : m) (j : n₁) : + (fromCols A₁ A₂) i (Sum.inl j) = A₁ i j := rfl + +@[deprecated (since := "2024-12-11")] alias fromColumns_apply_inl := fromCols_apply_inl @[simp] -lemma fromColumns_apply_inr (A₁ : Matrix m n₁ R) (A₂ : Matrix m n₂ R) (i : m) (j : n₂) : - (fromColumns A₁ A₂) i (Sum.inr j) = A₂ i j := rfl +lemma fromCols_apply_inr (A₁ : Matrix m n₁ R) (A₂ : Matrix m n₂ R) (i : m) (j : n₂) : + (fromCols A₁ A₂) i (Sum.inr j) = A₂ i j := rfl + +@[deprecated (since := "2024-12-11")] alias fromColumns_apply_inr := fromCols_apply_inr @[simp] lemma toRows₁_apply (A : Matrix (m₁ ⊕ m₂) n R) (i : m₁) (j : n) : @@ -81,26 +89,36 @@ lemma toRows₂_fromRows (A₁ : Matrix m₁ n R) (A₂ : Matrix m₂ n R) : toRows₂ (fromRows A₁ A₂) = A₂ := rfl @[simp] -lemma toColumns₁_apply (A : Matrix m (n₁ ⊕ n₂) R) (i : m) (j : n₁) : - (toColumns₁ A) i j = A i (Sum.inl j) := rfl +lemma toCols₁_apply (A : Matrix m (n₁ ⊕ n₂) R) (i : m) (j : n₁) : + (toCols₁ A) i j = A i (Sum.inl j) := rfl + +@[deprecated (since := "2024-12-11")] alias toColumns₁_apply := toCols₁_apply @[simp] -lemma toColumns₂_apply (A : Matrix m (n₁ ⊕ n₂) R) (i : m) (j : n₂) : - (toColumns₂ A) i j = A i (Sum.inr j) := rfl +lemma toCols₂_apply (A : Matrix m (n₁ ⊕ n₂) R) (i : m) (j : n₂) : + (toCols₂ A) i j = A i (Sum.inr j) := rfl + +@[deprecated (since := "2024-12-11")] alias toColumns₂_apply := toCols₂_apply @[simp] -lemma toColumns₁_fromColumns (A₁ : Matrix m n₁ R) (A₂ : Matrix m n₂ R) : - toColumns₁ (fromColumns A₁ A₂) = A₁ := rfl +lemma toCols₁_fromCols (A₁ : Matrix m n₁ R) (A₂ : Matrix m n₂ R) : + toCols₁ (fromCols A₁ A₂) = A₁ := rfl + +@[deprecated (since := "2024-12-11")] alias toColumns₁_fromColumns := toCols₁_fromCols @[simp] -lemma toColumns₂_fromColumns (A₁ : Matrix m n₁ R) (A₂ : Matrix m n₂ R) : - toColumns₂ (fromColumns A₁ A₂) = A₂ := rfl +lemma toCols₂_fromCols (A₁ : Matrix m n₁ R) (A₂ : Matrix m n₂ R) : + toCols₂ (fromCols A₁ A₂) = A₂ := rfl + +@[deprecated (since := "2024-12-11")] alias toColumns₂_fromColumns := toCols₂_fromCols @[simp] -lemma fromColumns_toColumns (A : Matrix m (n₁ ⊕ n₂) R) : - fromColumns A.toColumns₁ A.toColumns₂ = A := by +lemma fromCols_toCols (A : Matrix m (n₁ ⊕ n₂) R) : + fromCols A.toCols₁ A.toCols₂ = A := by ext i (j | j) <;> simp +@[deprecated (since := "2024-12-11")] alias fromColumns_toColumns := fromCols_toCols + @[simp] lemma fromRows_toRows (A : Matrix (m₁ ⊕ m₂) n R) : fromRows A.toRows₁ A.toRows₂ = A := by ext (i | i) j <;> simp @@ -110,14 +128,18 @@ lemma fromRows_inj : Function.Injective2 (@fromRows R m₁ m₂ n) := by simp only [funext_iff, ← Matrix.ext_iff] aesop -lemma fromColumns_inj : Function.Injective2 (@fromColumns R m n₁ n₂) := by +lemma fromCols_inj : Function.Injective2 (@fromCols R m n₁ n₂) := by intros x1 x2 y1 y2 simp only [funext_iff, ← Matrix.ext_iff] aesop -lemma fromColumns_ext_iff (A₁ : Matrix m n₁ R) (A₂ : Matrix m n₂ R) (B₁ : Matrix m n₁ R) +@[deprecated (since := "2024-12-11")] alias fromColumns_inj := fromCols_inj + +lemma fromCols_ext_iff (A₁ : Matrix m n₁ R) (A₂ : Matrix m n₂ R) (B₁ : Matrix m n₁ R) (B₂ : Matrix m n₂ R) : - fromColumns A₁ A₂ = fromColumns B₁ B₂ ↔ A₁ = B₁ ∧ A₂ = B₂ := fromColumns_inj.eq_iff + fromCols A₁ A₂ = fromCols B₁ B₂ ↔ A₁ = B₁ ∧ A₂ = B₂ := fromCols_inj.eq_iff + +@[deprecated (since := "2024-12-11")] alias fromColumns_ext_iff := fromCols_ext_iff lemma fromRows_ext_iff (A₁ : Matrix m₁ n R) (A₂ : Matrix m₂ n R) (B₁ : Matrix m₁ n R) (B₂ : Matrix m₂ n R) : @@ -125,14 +147,16 @@ lemma fromRows_ext_iff (A₁ : Matrix m₁ n R) (A₂ : Matrix m₂ n R) (B₁ : /-- A column partitioned matrix when transposed gives a row partitioned matrix with columns of the initial matrix transposed to become rows. -/ -lemma transpose_fromColumns (A₁ : Matrix m n₁ R) (A₂ : Matrix m n₂ R) : - transpose (fromColumns A₁ A₂) = fromRows (transpose A₁) (transpose A₂) := by +lemma transpose_fromCols (A₁ : Matrix m n₁ R) (A₂ : Matrix m n₂ R) : + transpose (fromCols A₁ A₂) = fromRows (transpose A₁) (transpose A₂) := by ext (i | i) j <;> simp +@[deprecated (since := "2024-12-11")] alias transpose_fromColumns := transpose_fromCols + /-- A row partitioned matrix when transposed gives a column partitioned matrix with rows of the initial matrix transposed to become columns. -/ lemma transpose_fromRows (A₁ : Matrix m₁ n R) (A₂ : Matrix m₂ n R) : - transpose (fromRows A₁ A₂) = fromColumns (transpose A₁) (transpose A₂) := by + transpose (fromRows A₁ A₂) = fromCols (transpose A₁) (transpose A₂) := by ext i (j | j) <;> simp section Neg @@ -147,24 +171,32 @@ lemma fromRows_neg (A₁ : Matrix m₁ n R) (A₂ : Matrix m₂ n R) : /-- Negating a matrix partitioned by columns is equivalent to negating each of the columns. -/ @[simp] -lemma fromColumns_neg (A₁ : Matrix n m₁ R) (A₂ : Matrix n m₂ R) : - -fromColumns A₁ A₂ = fromColumns (-A₁) (-A₂) := by +lemma fromCols_neg (A₁ : Matrix n m₁ R) (A₂ : Matrix n m₂ R) : + -fromCols A₁ A₂ = fromCols (-A₁) (-A₂) := by ext i (j | j) <;> simp +@[deprecated (since := "2024-12-11")] alias fromColumns_neg := fromCols_neg + end Neg @[simp] -lemma fromColumns_fromRows_eq_fromBlocks (B₁₁ : Matrix m₁ n₁ R) (B₁₂ : Matrix m₁ n₂ R) +lemma fromCols_fromRows_eq_fromBlocks (B₁₁ : Matrix m₁ n₁ R) (B₁₂ : Matrix m₁ n₂ R) (B₂₁ : Matrix m₂ n₁ R) (B₂₂ : Matrix m₂ n₂ R) : - fromColumns (fromRows B₁₁ B₂₁) (fromRows B₁₂ B₂₂) = fromBlocks B₁₁ B₁₂ B₂₁ B₂₂ := by + fromCols (fromRows B₁₁ B₂₁) (fromRows B₁₂ B₂₂) = fromBlocks B₁₁ B₁₂ B₂₁ B₂₂ := by ext (_ | _) (_ | _) <;> simp +@[deprecated (since := "2024-12-11")] +alias fromColumns_fromRows_eq_fromBlocks := fromCols_fromRows_eq_fromBlocks + @[simp] -lemma fromRows_fromColumn_eq_fromBlocks (B₁₁ : Matrix m₁ n₁ R) (B₁₂ : Matrix m₁ n₂ R) +lemma fromRows_fromCols_eq_fromBlocks (B₁₁ : Matrix m₁ n₁ R) (B₁₂ : Matrix m₁ n₂ R) (B₂₁ : Matrix m₂ n₁ R) (B₂₂ : Matrix m₂ n₂ R) : - fromRows (fromColumns B₁₁ B₁₂) (fromColumns B₂₁ B₂₂) = fromBlocks B₁₁ B₁₂ B₂₁ B₂₂ := by + fromRows (fromCols B₁₁ B₁₂) (fromCols B₂₁ B₂₂) = fromBlocks B₁₁ B₁₂ B₂₁ B₂₂ := by ext (_ | _) (_ | _) <;> simp +@[deprecated (since := "2024-12-11")] +alias fromRows_fromColumn_eq_fromBlocks := fromRows_fromCols_eq_fromBlocks + section Semiring variable [Semiring R] @@ -175,10 +207,12 @@ lemma fromRows_mulVec [Fintype n] (A₁ : Matrix m₁ n R) (A₂ : Matrix m₂ n ext (_ | _) <;> rfl @[simp] -lemma vecMul_fromColumns [Fintype m] (B₁ : Matrix m n₁ R) (B₂ : Matrix m n₂ R) (v : m → R) : - v ᵥ* fromColumns B₁ B₂ = Sum.elim (v ᵥ* B₁) (v ᵥ* B₂) := by +lemma vecMul_fromCols [Fintype m] (B₁ : Matrix m n₁ R) (B₂ : Matrix m n₂ R) (v : m → R) : + v ᵥ* fromCols B₁ B₂ = Sum.elim (v ᵥ* B₁) (v ᵥ* B₂) := by ext (_ | _) <;> rfl +@[deprecated (since := "2024-12-11")] alias vecMul_fromColumns := vecMul_fromCols + @[simp] lemma sum_elim_vecMul_fromRows [Fintype m₁] [Fintype m₂] (B₁ : Matrix m₁ n R) (B₂ : Matrix m₂ n R) (v₁ : m₁ → R) (v₂ : m₂ → R) : @@ -187,11 +221,13 @@ lemma sum_elim_vecMul_fromRows [Fintype m₁] [Fintype m₂] (B₁ : Matrix m₁ simp [Matrix.vecMul, fromRows, dotProduct] @[simp] -lemma fromColumns_mulVec_sum_elim [Fintype n₁] [Fintype n₂] +lemma fromCols_mulVec_sum_elim [Fintype n₁] [Fintype n₂] (A₁ : Matrix m n₁ R) (A₂ : Matrix m n₂ R) (v₁ : n₁ → R) (v₂ : n₂ → R) : - fromColumns A₁ A₂ *ᵥ Sum.elim v₁ v₂ = A₁ *ᵥ v₁ + A₂ *ᵥ v₂ := by + fromCols A₁ A₂ *ᵥ Sum.elim v₁ v₂ = A₁ *ᵥ v₁ + A₂ *ᵥ v₂ := by ext - simp [Matrix.mulVec, fromColumns] + simp [Matrix.mulVec, fromCols] + +@[deprecated (since := "2024-12-11")] alias fromColumns_mulVec_sum_elim := fromCols_mulVec_sum_elim @[simp] lemma fromRows_mul [Fintype n] (A₁ : Matrix m₁ n R) (A₂ : Matrix m₂ n R) (B : Matrix n m R) : @@ -199,42 +235,52 @@ lemma fromRows_mul [Fintype n] (A₁ : Matrix m₁ n R) (A₂ : Matrix m₂ n R) ext (_ | _) _ <;> simp [mul_apply] @[simp] -lemma mul_fromColumns [Fintype n] (A : Matrix m n R) (B₁ : Matrix n n₁ R) (B₂ : Matrix n n₂ R) : - A * fromColumns B₁ B₂ = fromColumns (A * B₁) (A * B₂) := by +lemma mul_fromCols [Fintype n] (A : Matrix m n R) (B₁ : Matrix n n₁ R) (B₂ : Matrix n n₂ R) : + A * fromCols B₁ B₂ = fromCols (A * B₁) (A * B₂) := by ext _ (_ | _) <;> simp [mul_apply] +@[deprecated (since := "2024-12-11")] alias mul_fromColumns := mul_fromCols + @[simp] lemma fromRows_zero : fromRows (0 : Matrix m₁ n R) (0 : Matrix m₂ n R) = 0 := by ext (_ | _) _ <;> simp @[simp] -lemma fromColumns_zero : fromColumns (0 : Matrix m n₁ R) (0 : Matrix m n₂ R) = 0 := by +lemma fromCols_zero : fromCols (0 : Matrix m n₁ R) (0 : Matrix m n₂ R) = 0 := by ext _ (_ | _) <;> simp +@[deprecated (since := "2024-12-11")] alias fromColumns_zero := fromCols_zero + /-- A row partitioned matrix multiplied by a column partitioned matrix gives a 2 by 2 block matrix. -/ -lemma fromRows_mul_fromColumns [Fintype n] (A₁ : Matrix m₁ n R) (A₂ : Matrix m₂ n R) +lemma fromRows_mul_fromCols [Fintype n] (A₁ : Matrix m₁ n R) (A₂ : Matrix m₂ n R) (B₁ : Matrix n n₁ R) (B₂ : Matrix n n₂ R) : - (fromRows A₁ A₂) * (fromColumns B₁ B₂) = + (fromRows A₁ A₂) * (fromCols B₁ B₂) = fromBlocks (A₁ * B₁) (A₁ * B₂) (A₂ * B₁) (A₂ * B₂) := by ext (_ | _) (_ | _) <;> simp +@[deprecated (since := "2024-12-11")] alias fromRows_mul_fromColumns := fromRows_mul_fromCols + /-- A column partitioned matrix multiplied by a row partitioned matrix gives the sum of the "outer" products of the block matrices. -/ -lemma fromColumns_mul_fromRows [Fintype n₁] [Fintype n₂] (A₁ : Matrix m n₁ R) (A₂ : Matrix m n₂ R) +lemma fromCols_mul_fromRows [Fintype n₁] [Fintype n₂] (A₁ : Matrix m n₁ R) (A₂ : Matrix m n₂ R) (B₁ : Matrix n₁ n R) (B₂ : Matrix n₂ n R) : - fromColumns A₁ A₂ * fromRows B₁ B₂ = (A₁ * B₁ + A₂ * B₂) := by + fromCols A₁ A₂ * fromRows B₁ B₂ = (A₁ * B₁ + A₂ * B₂) := by ext simp [mul_apply] +@[deprecated (since := "2024-12-11")] alias fromColumns_mul_fromRows := fromCols_mul_fromRows + /-- A column partitioned matrix multipiled by a block matrix results in a column partitioned matrix. -/ -lemma fromColumns_mul_fromBlocks [Fintype m₁] [Fintype m₂] (A₁ : Matrix m m₁ R) (A₂ : Matrix m m₂ R) +lemma fromCols_mul_fromBlocks [Fintype m₁] [Fintype m₂] (A₁ : Matrix m m₁ R) (A₂ : Matrix m m₂ R) (B₁₁ : Matrix m₁ n₁ R) (B₁₂ : Matrix m₁ n₂ R) (B₂₁ : Matrix m₂ n₁ R) (B₂₂ : Matrix m₂ n₂ R) : - (fromColumns A₁ A₂) * fromBlocks B₁₁ B₁₂ B₂₁ B₂₂ = - fromColumns (A₁ * B₁₁ + A₂ * B₂₁) (A₁ * B₁₂ + A₂ * B₂₂) := by + (fromCols A₁ A₂) * fromBlocks B₁₁ B₁₂ B₂₁ B₂₂ = + fromCols (A₁ * B₁₁ + A₂ * B₂₁) (A₁ * B₁₂ + A₂ * B₂₂) := by ext _ (_ | _) <;> simp [mul_apply] +@[deprecated (since := "2024-12-11")] alias fromColumns_mul_fromBlocks := fromCols_mul_fromBlocks + /-- A block matrix multiplied by a row partitioned matrix gives a row partitioned matrix. -/ lemma fromBlocks_mul_fromRows [Fintype n₁] [Fintype n₂] (A₁ : Matrix n₁ n R) (A₂ : Matrix n₂ n R) (B₁₁ : Matrix m₁ n₁ R) (B₁₂ : Matrix m₁ n₂ R) (B₂₁ : Matrix m₂ n₁ R) (B₂₂ : Matrix m₂ n₂ R) : @@ -251,23 +297,30 @@ variable [CommRing R] /-- Multiplication of a matrix by its inverse is commutative. This is the column and row partitioned matrix form of `Matrix.mul_eq_one_comm`. -The condition `e : n ≃ n₁ ⊕ n₂` states that `fromColumns A₁ A₂` and `fromRows B₁ B₂` are "square". +The condition `e : n ≃ n₁ ⊕ n₂` states that `fromCols A₁ A₂` and `fromRows B₁ B₂` are "square". -/ -lemma fromColumns_mul_fromRows_eq_one_comm +lemma fromCols_mul_fromRows_eq_one_comm [Fintype n₁] [Fintype n₂] [Fintype n] [DecidableEq n] [DecidableEq n₁] [DecidableEq n₂] (e : n ≃ n₁ ⊕ n₂) (A₁ : Matrix n n₁ R) (A₂ : Matrix n n₂ R) (B₁ : Matrix n₁ n R) (B₂ : Matrix n₂ n R) : - fromColumns A₁ A₂ * fromRows B₁ B₂ = 1 ↔ fromRows B₁ B₂ * fromColumns A₁ A₂ = 1 := + fromCols A₁ A₂ * fromRows B₁ B₂ = 1 ↔ fromRows B₁ B₂ * fromCols A₁ A₂ = 1 := mul_eq_one_comm_of_equiv e -/-- The lemma `fromColumns_mul_fromRows_eq_one_comm` specialized to the case where the index sets n₁ -and n₂, are the result of subtyping by a predicate and its complement. -/ -lemma equiv_compl_fromColumns_mul_fromRows_eq_one_comm +@[deprecated (since := "2024-12-11")] +alias fromColumns_mul_fromRows_eq_one_comm := fromCols_mul_fromRows_eq_one_comm + +/-- The lemma `fromCols_mul_fromRows_eq_one_comm` specialized to the case where the index sets +`n₁` and `n₂`, are the result of subtyping by a predicate and its complement. -/ +lemma equiv_compl_fromCols_mul_fromRows_eq_one_comm [Fintype n] [DecidableEq n] (p : n → Prop) [DecidablePred p] (A₁ : Matrix n {i // p i} R) (A₂ : Matrix n {i // ¬p i} R) (B₁ : Matrix {i // p i} n R) (B₂ : Matrix {i // ¬p i} n R) : - fromColumns A₁ A₂ * fromRows B₁ B₂ = 1 ↔ fromRows B₁ B₂ * fromColumns A₁ A₂ = 1 := - fromColumns_mul_fromRows_eq_one_comm (id (Equiv.sumCompl p).symm) A₁ A₂ B₁ B₂ + fromCols A₁ A₂ * fromRows B₁ B₂ = 1 ↔ fromRows B₁ B₂ * fromCols A₁ A₂ = 1 := + fromCols_mul_fromRows_eq_one_comm (id (Equiv.sumCompl p).symm) A₁ A₂ B₁ B₂ + +@[deprecated (since := "2024-12-11")] +alias equiv_compl_fromColumns_mul_fromRows_eq_one_comm := + equiv_compl_fromCols_mul_fromRows_eq_one_comm end CommRing @@ -276,18 +329,26 @@ variable [Star R] /-- A column partitioned matrix in a Star ring when conjugate transposed gives a row partitioned matrix with the columns of the initial matrix conjugate transposed to become rows. -/ -lemma conjTranspose_fromColumns_eq_fromRows_conjTranspose (A₁ : Matrix m n₁ R) +lemma conjTranspose_fromCols_eq_fromRows_conjTranspose (A₁ : Matrix m n₁ R) (A₂ : Matrix m n₂ R) : - conjTranspose (fromColumns A₁ A₂) = fromRows (conjTranspose A₁) (conjTranspose A₂) := by + conjTranspose (fromCols A₁ A₂) = fromRows (conjTranspose A₁) (conjTranspose A₂) := by ext (_ | _) _ <;> simp +@[deprecated (since := "2024-12-11")] +alias conjTranspose_fromColumns_eq_fromRows_conjTranspose := + conjTranspose_fromCols_eq_fromRows_conjTranspose + /-- A row partitioned matrix in a Star ring when conjugate transposed gives a column partitioned matrix with the rows of the initial matrix conjugate transposed to become columns. -/ -lemma conjTranspose_fromRows_eq_fromColumns_conjTranspose (A₁ : Matrix m₁ n R) +lemma conjTranspose_fromRows_eq_fromCols_conjTranspose (A₁ : Matrix m₁ n R) (A₂ : Matrix m₂ n R) : conjTranspose (fromRows A₁ A₂) = - fromColumns (conjTranspose A₁) (conjTranspose A₂) := by + fromCols (conjTranspose A₁) (conjTranspose A₂) := by ext _ (_ | _) <;> simp +@[deprecated (since := "2024-12-11")] +alias conjTranspose_fromRows_eq_fromColumns_conjTranspose := + conjTranspose_fromRows_eq_fromCols_conjTranspose + end Star end Matrix diff --git a/Mathlib/Data/Matrix/Notation.lean b/Mathlib/Data/Matrix/Notation.lean index 6d9f367504b17..a150ccd3763b4 100644 --- a/Mathlib/Data/Matrix/Notation.lean +++ b/Mathlib/Data/Matrix/Notation.lean @@ -381,9 +381,12 @@ theorem submatrix_updateRow_succAbove (A : Matrix (Fin m.succ) n' α) (v : n' /-- Updating a column then removing it is the same as removing it. -/ @[simp] -theorem submatrix_updateColumn_succAbove (A : Matrix m' (Fin n.succ) α) (v : m' → α) (f : o' → m') - (i : Fin n.succ) : (A.updateColumn i v).submatrix f i.succAbove = A.submatrix f i.succAbove := - ext fun _r s => updateColumn_ne (Fin.succAbove_ne i s) +theorem submatrix_updateCol_succAbove (A : Matrix m' (Fin n.succ) α) (v : m' → α) (f : o' → m') + (i : Fin n.succ) : (A.updateCol i v).submatrix f i.succAbove = A.submatrix f i.succAbove := + ext fun _r s => updateCol_ne (Fin.succAbove_ne i s) + +@[deprecated (since := "2024-12-11")] +alias submatrix_updateColumn_succAbove := submatrix_updateCol_succAbove end Submatrix diff --git a/Mathlib/Data/Matrix/RowCol.lean b/Mathlib/Data/Matrix/RowCol.lean index 43c2b513fbabd..562dd28052b82 100644 --- a/Mathlib/Data/Matrix/RowCol.lean +++ b/Mathlib/Data/Matrix/RowCol.lean @@ -171,9 +171,11 @@ def updateRow [DecidableEq m] (M : Matrix m n α) (i : m) (b : n → α) : Matri of <| Function.update M i b /-- Update, i.e. replace the `j`th column of matrix `A` with the values in `b`. -/ -def updateColumn [DecidableEq n] (M : Matrix m n α) (j : n) (b : m → α) : Matrix m n α := +def updateCol [DecidableEq n] (M : Matrix m n α) (j : n) (b : m → α) : Matrix m n α := of fun i => Function.update (M i) j (b i) +@[deprecated (since := "2024-12-11")] alias updateColumn := updateCol + variable {M : Matrix m n α} {i : m} {j : n} {b : n → α} {c : m → α} @[simp] @@ -182,44 +184,52 @@ theorem updateRow_self [DecidableEq m] : updateRow M i b i = b := Function.update_same (β := fun _ => (n → α)) i b M @[simp] -theorem updateColumn_self [DecidableEq n] : updateColumn M j c i j = c i := +theorem updateCol_self [DecidableEq n] : updateCol M j c i j = c i := -- Porting note: (implicit arg) added `(β := _)` Function.update_same (β := fun _ => α) j (c i) (M i) +@[deprecated (since := "2024-12-11")] alias updateColumn_self := updateCol_self + @[simp] theorem updateRow_ne [DecidableEq m] {i' : m} (i_ne : i' ≠ i) : updateRow M i b i' = M i' := -- Porting note: (implicit arg) added `(β := _)` Function.update_noteq (β := fun _ => (n → α)) i_ne b M @[simp] -theorem updateColumn_ne [DecidableEq n] {j' : n} (j_ne : j' ≠ j) : - updateColumn M j c i j' = M i j' := +theorem updateCol_ne [DecidableEq n] {j' : n} (j_ne : j' ≠ j) : + updateCol M j c i j' = M i j' := -- Porting note: (implicit arg) added `(β := _)` Function.update_noteq (β := fun _ => α) j_ne (c i) (M i) +@[deprecated (since := "2024-12-11")] alias updateColumn_ne := updateCol_ne + theorem updateRow_apply [DecidableEq m] {i' : m} : updateRow M i b i' j = if i' = i then b j else M i' j := by by_cases h : i' = i · rw [h, updateRow_self, if_pos rfl] · rw [updateRow_ne h, if_neg h] -theorem updateColumn_apply [DecidableEq n] {j' : n} : - updateColumn M j c i j' = if j' = j then c i else M i j' := by +theorem updateCol_apply [DecidableEq n] {j' : n} : + updateCol M j c i j' = if j' = j then c i else M i j' := by by_cases h : j' = j - · rw [h, updateColumn_self, if_pos rfl] - · rw [updateColumn_ne h, if_neg h] + · rw [h, updateCol_self, if_pos rfl] + · rw [updateCol_ne h, if_neg h] + +@[deprecated (since := "2024-12-11")] alias updateColumn_apply := updateCol_apply @[simp] -theorem updateColumn_subsingleton [Subsingleton n] (A : Matrix m n R) (i : n) (b : m → R) : - A.updateColumn i b = (col (Fin 1) b).submatrix id (Function.const n 0) := by +theorem updateCol_subsingleton [Subsingleton n] (A : Matrix m n R) (i : n) (b : m → R) : + A.updateCol i b = (col (Fin 1) b).submatrix id (Function.const n 0) := by ext x y - simp [updateColumn_apply, Subsingleton.elim i y] + simp [updateCol_apply, Subsingleton.elim i y] + +@[deprecated (since := "2024-12-11")] alias updateColumn_subsingleton := updateCol_subsingleton @[simp] theorem updateRow_subsingleton [Subsingleton m] (A : Matrix m n R) (i : m) (b : n → R) : A.updateRow i b = (row (Fin 1) b).submatrix (Function.const m 0) id := by ext x y - simp [updateColumn_apply, Subsingleton.elim i x] + simp [updateCol_apply, Subsingleton.elim i x] theorem map_updateRow [DecidableEq m] (f : α → β) : map (updateRow M i b) f = updateRow (M.map f) i (f ∘ b) := by @@ -227,59 +237,70 @@ theorem map_updateRow [DecidableEq m] (f : α → β) : rw [updateRow_apply, map_apply, map_apply, updateRow_apply] exact apply_ite f _ _ _ -theorem map_updateColumn [DecidableEq n] (f : α → β) : - map (updateColumn M j c) f = updateColumn (M.map f) j (f ∘ c) := by +theorem map_updateCol [DecidableEq n] (f : α → β) : + map (updateCol M j c) f = updateCol (M.map f) j (f ∘ c) := by ext - rw [updateColumn_apply, map_apply, map_apply, updateColumn_apply] + rw [updateCol_apply, map_apply, map_apply, updateCol_apply] exact apply_ite f _ _ _ -theorem updateRow_transpose [DecidableEq n] : updateRow Mᵀ j c = (updateColumn M j c)ᵀ := by +@[deprecated (since := "2024-12-11")] alias map_updateColumn := map_updateCol + +theorem updateRow_transpose [DecidableEq n] : updateRow Mᵀ j c = (updateCol M j c)ᵀ := by ext - rw [transpose_apply, updateRow_apply, updateColumn_apply] + rw [transpose_apply, updateRow_apply, updateCol_apply] rfl -theorem updateColumn_transpose [DecidableEq m] : updateColumn Mᵀ i b = (updateRow M i b)ᵀ := by +theorem updateCol_transpose [DecidableEq m] : updateCol Mᵀ i b = (updateRow M i b)ᵀ := by ext - rw [transpose_apply, updateRow_apply, updateColumn_apply] + rw [transpose_apply, updateRow_apply, updateCol_apply] rfl +@[deprecated (since := "2024-12-11")] alias updateColumn_transpose := updateCol_transpose + theorem updateRow_conjTranspose [DecidableEq n] [Star α] : - updateRow Mᴴ j (star c) = (updateColumn M j c)ᴴ := by + updateRow Mᴴ j (star c) = (updateCol M j c)ᴴ := by rw [conjTranspose, conjTranspose, transpose_map, transpose_map, updateRow_transpose, - map_updateColumn] + map_updateCol] rfl -theorem updateColumn_conjTranspose [DecidableEq m] [Star α] : - updateColumn Mᴴ i (star b) = (updateRow M i b)ᴴ := by - rw [conjTranspose, conjTranspose, transpose_map, transpose_map, updateColumn_transpose, +theorem updateCol_conjTranspose [DecidableEq m] [Star α] : + updateCol Mᴴ i (star b) = (updateRow M i b)ᴴ := by + rw [conjTranspose, conjTranspose, transpose_map, transpose_map, updateCol_transpose, map_updateRow] rfl +@[deprecated (since := "2024-12-11")] alias updateColumn_conjTranspose := updateCol_conjTranspose + @[simp] theorem updateRow_eq_self [DecidableEq m] (A : Matrix m n α) (i : m) : A.updateRow i (A i) = A := Function.update_eq_self i A @[simp] -theorem updateColumn_eq_self [DecidableEq n] (A : Matrix m n α) (i : n) : - (A.updateColumn i fun j => A j i) = A := +theorem updateCol_eq_self [DecidableEq n] (A : Matrix m n α) (i : n) : + (A.updateCol i fun j => A j i) = A := funext fun j => Function.update_eq_self i (A j) -theorem diagonal_updateColumn_single [DecidableEq n] [Zero α] (v : n → α) (i : n) (x : α) : - (diagonal v).updateColumn i (Pi.single i x) = diagonal (Function.update v i x) := by +@[deprecated (since := "2024-12-11")] alias updateColumn_eq_self := updateCol_eq_self + +theorem diagonal_updateCol_single [DecidableEq n] [Zero α] (v : n → α) (i : n) (x : α) : + (diagonal v).updateCol i (Pi.single i x) = diagonal (Function.update v i x) := by ext j k obtain rfl | hjk := eq_or_ne j k · rw [diagonal_apply_eq] obtain rfl | hji := eq_or_ne j i - · rw [updateColumn_self, Pi.single_eq_same, Function.update_same] - · rw [updateColumn_ne hji, diagonal_apply_eq, Function.update_noteq hji] + · rw [updateCol_self, Pi.single_eq_same, Function.update_same] + · rw [updateCol_ne hji, diagonal_apply_eq, Function.update_noteq hji] · rw [diagonal_apply_ne _ hjk] obtain rfl | hki := eq_or_ne k i - · rw [updateColumn_self, Pi.single_eq_of_ne hjk] - · rw [updateColumn_ne hki, diagonal_apply_ne _ hjk] + · rw [updateCol_self, Pi.single_eq_of_ne hjk] + · rw [updateCol_ne hki, diagonal_apply_ne _ hjk] + +@[deprecated (since := "2024-12-11")] +alias diagonal_updateColumn_single := diagonal_updateCol_single theorem diagonal_updateRow_single [DecidableEq n] [Zero α] (v : n → α) (i : n) (x : α) : (diagonal v).updateRow i (Pi.single i x) = diagonal (Function.update v i x) := by - rw [← diagonal_transpose, updateRow_transpose, diagonal_updateColumn_single, diagonal_transpose] + rw [← diagonal_transpose, updateRow_transpose, diagonal_updateCol_single, diagonal_transpose] /-! Updating rows and columns commutes in the obvious way with reindexing the matrix. -/ @@ -295,16 +316,22 @@ theorem submatrix_updateRow_equiv [DecidableEq l] [DecidableEq m] (A : Matrix m (A.updateRow i r).submatrix e f = updateRow (A.submatrix e f) (e.symm i) fun i => r (f i) := Eq.trans (by simp_rw [Equiv.apply_symm_apply]) (updateRow_submatrix_equiv A _ _ e f).symm -theorem updateColumn_submatrix_equiv [DecidableEq o] [DecidableEq n] (A : Matrix m n α) (j : o) - (c : l → α) (e : l ≃ m) (f : o ≃ n) : updateColumn (A.submatrix e f) j c = - (A.updateColumn (f j) fun i => c (e.symm i)).submatrix e f := by +theorem updateCol_submatrix_equiv [DecidableEq o] [DecidableEq n] (A : Matrix m n α) (j : o) + (c : l → α) (e : l ≃ m) (f : o ≃ n) : updateCol (A.submatrix e f) j c = + (A.updateCol (f j) fun i => c (e.symm i)).submatrix e f := by simpa only [← transpose_submatrix, updateRow_transpose] using congr_arg transpose (updateRow_submatrix_equiv Aᵀ j c f e) -theorem submatrix_updateColumn_equiv [DecidableEq o] [DecidableEq n] (A : Matrix m n α) (j : n) - (c : m → α) (e : l ≃ m) (f : o ≃ n) : (A.updateColumn j c).submatrix e f = - updateColumn (A.submatrix e f) (f.symm j) fun i => c (e i) := - Eq.trans (by simp_rw [Equiv.apply_symm_apply]) (updateColumn_submatrix_equiv A _ _ e f).symm +@[deprecated (since := "2024-12-11")] +alias updateColumn_submatrix_equiv := updateCol_submatrix_equiv + +theorem submatrix_updateCol_equiv [DecidableEq o] [DecidableEq n] (A : Matrix m n α) (j : n) + (c : m → α) (e : l ≃ m) (f : o ≃ n) : (A.updateCol j c).submatrix e f = + updateCol (A.submatrix e f) (f.symm j) fun i => c (e i) := + Eq.trans (by simp_rw [Equiv.apply_symm_apply]) (updateCol_submatrix_equiv A _ _ e f).symm + +@[deprecated (since := "2024-12-11")] +alias submatrix_updateColumn_equiv := submatrix_updateCol_equiv /-! `reindex` versions of the above `submatrix` lemmas for convenience. -/ @@ -319,14 +346,18 @@ theorem reindex_updateRow [DecidableEq l] [DecidableEq m] (A : Matrix m n α) (i reindex e f (A.updateRow i r) = updateRow (reindex e f A) (e i) fun i => r (f.symm i) := submatrix_updateRow_equiv _ _ _ _ _ -theorem updateColumn_reindex [DecidableEq o] [DecidableEq n] (A : Matrix m n α) (j : o) (c : l → α) +theorem updateCol_reindex [DecidableEq o] [DecidableEq n] (A : Matrix m n α) (j : o) (c : l → α) (e : m ≃ l) (f : n ≃ o) : - updateColumn (reindex e f A) j c = reindex e f (A.updateColumn (f.symm j) fun i => c (e i)) := - updateColumn_submatrix_equiv _ _ _ _ _ + updateCol (reindex e f A) j c = reindex e f (A.updateCol (f.symm j) fun i => c (e i)) := + updateCol_submatrix_equiv _ _ _ _ _ + +@[deprecated (since := "2024-12-11")] alias updateColumn_reindex := updateCol_reindex -theorem reindex_updateColumn [DecidableEq o] [DecidableEq n] (A : Matrix m n α) (j : n) (c : m → α) +theorem reindex_updateCol [DecidableEq o] [DecidableEq n] (A : Matrix m n α) (j : n) (c : m → α) (e : m ≃ l) (f : n ≃ o) : - reindex e f (A.updateColumn j c) = updateColumn (reindex e f A) (f j) fun i => c (e.symm i) := - submatrix_updateColumn_equiv _ _ _ _ _ + reindex e f (A.updateCol j c) = updateCol (reindex e f A) (f j) fun i => c (e.symm i) := + submatrix_updateCol_equiv _ _ _ _ _ + +@[deprecated (since := "2024-12-11")] alias reindex_updateColumn := reindex_updateCol end Matrix diff --git a/Mathlib/LinearAlgebra/Determinant.lean b/Mathlib/LinearAlgebra/Determinant.lean index dd5c969809a49..66160703d9068 100644 --- a/Mathlib/LinearAlgebra/Determinant.lean +++ b/Mathlib/LinearAlgebra/Determinant.lean @@ -455,13 +455,13 @@ nonrec def Basis.det : M [⋀^ι]→ₗ[R] R where cases Subsingleton.elim inst ‹_› simp only [e.toMatrix_update, LinearEquiv.map_add, Finsupp.coe_add] -- Porting note: was `exact det_update_column_add _ _ _ _` - convert det_updateColumn_add (e.toMatrix v) i (e.repr x) (e.repr y) + convert det_updateCol_add (e.toMatrix v) i (e.repr x) (e.repr y) map_update_smul' := by intro inst u i c x cases Subsingleton.elim inst ‹_› simp only [e.toMatrix_update, Algebra.id.smul_eq_mul, LinearEquiv.map_smul] -- Porting note: was `apply det_update_column_smul` - convert det_updateColumn_smul (e.toMatrix u) i c (e.repr x) + convert det_updateCol_smul (e.toMatrix u) i c (e.repr x) map_eq_zero_of_eq' := by intro v i j h hij -- Porting note: added diff --git a/Mathlib/LinearAlgebra/Matrix/Adjugate.lean b/Mathlib/LinearAlgebra/Matrix/Adjugate.lean index 9037e4a5083b9..44818b76d54be 100644 --- a/Mathlib/LinearAlgebra/Matrix/Adjugate.lean +++ b/Mathlib/LinearAlgebra/Matrix/Adjugate.lean @@ -70,11 +70,11 @@ variable (A : Matrix n n α) (b : n → α) Otherwise, the outcome of `cramerMap` is well-defined but not necessarily useful. -/ def cramerMap (i : n) : α := - (A.updateColumn i b).det + (A.updateCol i b).det theorem cramerMap_is_linear (i : n) : IsLinearMap α fun b => cramerMap A b i := - { map_add := det_updateColumn_add _ _ - map_smul := det_updateColumn_smul _ _ } + { map_add := det_updateCol_add _ _ + map_smul := det_updateCol_smul _ _ } theorem cramer_is_linear : IsLinearMap α (cramerMap A) := by constructor <;> intros <;> ext i @@ -90,11 +90,11 @@ theorem cramer_is_linear : IsLinearMap α (cramerMap A) := by def cramer (A : Matrix n n α) : (n → α) →ₗ[α] (n → α) := IsLinearMap.mk' (cramerMap A) (cramer_is_linear A) -theorem cramer_apply (i : n) : cramer A b i = (A.updateColumn i b).det := +theorem cramer_apply (i : n) : cramer A b i = (A.updateCol i b).det := rfl theorem cramer_transpose_apply (i : n) : cramer Aᵀ b i = (A.updateRow i b).det := by - rw [cramer_apply, updateColumn_transpose, det_transpose] + rw [cramer_apply, updateCol_transpose, det_transpose] theorem cramer_transpose_row_self (i : n) : Aᵀ.cramer (A i) = Pi.single i A.det := by ext j @@ -102,9 +102,9 @@ theorem cramer_transpose_row_self (i : n) : Aᵀ.cramer (A i) = Pi.single i A.de split_ifs with h · -- i = j: this entry should be `A.det` subst h - simp only [updateColumn_transpose, det_transpose, updateRow_eq_self] + simp only [updateCol_transpose, det_transpose, updateRow_eq_self] · -- i ≠ j: this entry should be 0 - rw [updateColumn_transpose, det_transpose] + rw [updateCol_transpose, det_transpose] apply det_zero_of_row_eq h rw [updateRow_self, updateRow_ne (Ne.symm h)] @@ -123,18 +123,18 @@ theorem cramer_one : cramer (1 : Matrix n n α) = 1 := by theorem cramer_smul (r : α) (A : Matrix n n α) : cramer (r • A) = r ^ (Fintype.card n - 1) • cramer A := - LinearMap.ext fun _ => funext fun _ => det_updateColumn_smul_left _ _ _ _ + LinearMap.ext fun _ => funext fun _ => det_updateCol_smul_left _ _ _ _ @[simp] theorem cramer_subsingleton_apply [Subsingleton n] (A : Matrix n n α) (b : n → α) (i : n) : - cramer A b i = b i := by rw [cramer_apply, det_eq_elem_of_subsingleton _ i, updateColumn_self] + cramer A b i = b i := by rw [cramer_apply, det_eq_elem_of_subsingleton _ i, updateCol_self] theorem cramer_zero [Nontrivial n] : cramer (0 : Matrix n n α) = 0 := by ext i j obtain ⟨j', hj'⟩ : ∃ j', j' ≠ j := exists_ne j apply det_eq_zero_of_column_eq_zero j' intro j'' - simp [updateColumn_ne hj'] + simp [updateCol_ne hj'] /-- Use linearity of `cramer` to take it out of a summation. -/ theorem sum_cramer {β} (s : Finset β) (f : β → n → α) : @@ -149,7 +149,7 @@ theorem sum_cramer_apply {β} (s : Finset β) (f : n → β → α) (i : n) : (Finset.sum_apply i s _).symm _ = cramer A (fun j : n => ∑ x ∈ s, f j x) i := by rw [sum_cramer, cramer_apply, cramer_apply] - simp only [updateColumn] + simp only [updateCol] congr with j congr apply Finset.sum_apply @@ -157,7 +157,7 @@ theorem sum_cramer_apply {β} (s : Finset β) (f : n → β → α) (i : n) : theorem cramer_submatrix_equiv (A : Matrix m m α) (e : n ≃ m) (b : n → α) : cramer (A.submatrix e e) b = cramer A (b ∘ e.symm) ∘ e := by ext i - simp_rw [Function.comp_apply, cramer_apply, updateColumn_submatrix_equiv, + simp_rw [Function.comp_apply, cramer_apply, updateCol_submatrix_equiv, det_submatrix_equiv_self e, Function.comp_def] theorem cramer_reindex (e : m ≃ n) (A : Matrix m m α) (b : n → α) : @@ -192,7 +192,7 @@ theorem adjugate_def (A : Matrix n n α) : adjugate A = of fun i => cramer Aᵀ theorem adjugate_apply (A : Matrix n n α) (i j : n) : adjugate A i j = (A.updateRow j (Pi.single i 1)).det := by - rw [adjugate_def, of_apply, cramer_apply, updateColumn_transpose, det_transpose] + rw [adjugate_def, of_apply, cramer_apply, updateCol_transpose, det_transpose] theorem adjugate_transpose (A : Matrix n n α) : (adjugate A)ᵀ = adjugate Aᵀ := by ext i j @@ -207,15 +207,15 @@ theorem adjugate_transpose (A : Matrix n n α) : (adjugate A)ᵀ = adjugate Aᵀ ext j' subst h have : σ j' = σ j ↔ j' = j := σ.injective.eq_iff - rw [updateRow_apply, updateColumn_apply] + rw [updateRow_apply, updateCol_apply] simp_rw [this] rw [← dite_eq_ite, ← dite_eq_ite] congr 1 with rfl rw [Pi.single_eq_same, Pi.single_eq_same] · -- Otherwise, we need to show that there is a `0` somewhere in the product. - have : (∏ j' : n, updateColumn A j (Pi.single i 1) (σ j') j') = 0 := by + have : (∏ j' : n, updateCol A j (Pi.single i 1) (σ j') j') = 0 := by apply prod_eq_zero (mem_univ j) - rw [updateColumn_self, Pi.single_eq_of_ne' h] + rw [updateCol_self, Pi.single_eq_of_ne' h] rw [this] apply prod_eq_zero (mem_univ (σ⁻¹ i)) erw [apply_symm_apply σ i, updateRow_self] @@ -297,7 +297,7 @@ theorem adjugate_zero [Nontrivial n] : adjugate (0 : Matrix n n α) = 0 := by obtain ⟨j', hj'⟩ : ∃ j', j' ≠ j := exists_ne j apply det_eq_zero_of_column_eq_zero j' intro j'' - simp [updateColumn_ne hj'] + simp [updateCol_ne hj'] @[simp] theorem adjugate_one : adjugate (1 : Matrix n n α) = 1 := by @@ -310,13 +310,13 @@ theorem adjugate_diagonal (v : n → α) : ext i j simp only [adjugate_def, cramer_apply, diagonal_transpose, of_apply] obtain rfl | hij := eq_or_ne i j - · rw [diagonal_apply_eq, diagonal_updateColumn_single, det_diagonal, + · rw [diagonal_apply_eq, diagonal_updateCol_single, det_diagonal, prod_update_of_mem (Finset.mem_univ _), sdiff_singleton_eq_erase, one_mul] · rw [diagonal_apply_ne _ hij] refine det_eq_zero_of_row_eq_zero j fun k => ?_ obtain rfl | hjk := eq_or_ne k j - · rw [updateColumn_self, Pi.single_eq_of_ne' hij] - · rw [updateColumn_ne hjk, diagonal_apply_ne' _ hjk] + · rw [updateCol_self, Pi.single_eq_of_ne' hij] + · rw [updateCol_ne hjk, diagonal_apply_ne' _ hjk] theorem _root_.RingHom.map_adjugate {R S : Type*} [CommRing R] [CommRing S] (f : R →+* S) (M : Matrix n n R) : f.mapMatrix M.adjugate = Matrix.adjugate (f.mapMatrix M) := by diff --git a/Mathlib/LinearAlgebra/Matrix/Basis.lean b/Mathlib/LinearAlgebra/Matrix/Basis.lean index f06b9e9d99955..f7bd16c3469f5 100644 --- a/Mathlib/LinearAlgebra/Matrix/Basis.lean +++ b/Mathlib/LinearAlgebra/Matrix/Basis.lean @@ -76,9 +76,9 @@ theorem toMatrix_self [DecidableEq ι] : e.toMatrix e = 1 := by simp [Basis.equivFun, Matrix.one_apply, Finsupp.single_apply, eq_comm] theorem toMatrix_update [DecidableEq ι'] (x : M) : - e.toMatrix (Function.update v j x) = Matrix.updateColumn (e.toMatrix v) j (e.repr x) := by + e.toMatrix (Function.update v j x) = Matrix.updateCol (e.toMatrix v) j (e.repr x) := by ext i' k - rw [Basis.toMatrix, Matrix.updateColumn_apply, e.toMatrix_apply] + rw [Basis.toMatrix, Matrix.updateCol_apply, e.toMatrix_apply] split_ifs with h · rw [h, update_same j x v] · rw [update_noteq h] diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean index 1577138e86598..274f69211c711 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean @@ -349,8 +349,10 @@ theorem det_updateRow_eq_zero (h : i ≠ j) : (M.updateRow j (M i)).det = 0 := det_zero_of_row_eq h (by simp [h]) /-- If we repeat a column of a matrix, we get a matrix of determinant zero. -/ -theorem det_updateColumn_eq_zero (h : i ≠ j) : - (M.updateColumn j (fun k ↦ M k i)).det = 0 := det_zero_of_column_eq h (by simp [h]) +theorem det_updateCol_eq_zero (h : i ≠ j) : + (M.updateCol j (fun k ↦ M k i)).det = 0 := det_zero_of_column_eq h (by simp [h]) + +@[deprecated (since := "2024-12-11")] alias det_updateColumn_eq_zero := det_updateCol_eq_zero end DetZero @@ -358,32 +360,37 @@ theorem det_updateRow_add (M : Matrix n n R) (j : n) (u v : n → R) : det (updateRow M j <| u + v) = det (updateRow M j u) + det (updateRow M j v) := (detRowAlternating : (n → R) [⋀^n]→ₗ[R] R).map_update_add M j u v -theorem det_updateColumn_add (M : Matrix n n R) (j : n) (u v : n → R) : - det (updateColumn M j <| u + v) = det (updateColumn M j u) + det (updateColumn M j v) := by +theorem det_updateCol_add (M : Matrix n n R) (j : n) (u v : n → R) : + det (updateCol M j <| u + v) = det (updateCol M j u) + det (updateCol M j v) := by rw [← det_transpose, ← updateRow_transpose, det_updateRow_add] simp [updateRow_transpose, det_transpose] +@[deprecated (since := "2024-12-11")] alias det_updateColumn_add := det_updateCol_add + theorem det_updateRow_smul (M : Matrix n n R) (j : n) (s : R) (u : n → R) : det (updateRow M j <| s • u) = s * det (updateRow M j u) := (detRowAlternating : (n → R) [⋀^n]→ₗ[R] R).map_update_smul M j s u -theorem det_updateColumn_smul (M : Matrix n n R) (j : n) (s : R) (u : n → R) : - det (updateColumn M j <| s • u) = s * det (updateColumn M j u) := by +theorem det_updateCol_smul (M : Matrix n n R) (j : n) (s : R) (u : n → R) : + det (updateCol M j <| s • u) = s * det (updateCol M j u) := by rw [← det_transpose, ← updateRow_transpose, det_updateRow_smul] simp [updateRow_transpose, det_transpose] +@[deprecated (since := "2024-12-11")] alias det_updateColumn_smul := det_updateCol_smul + theorem det_updateRow_smul_left (M : Matrix n n R) (j : n) (s : R) (u : n → R) : det (updateRow (s • M) j u) = s ^ (Fintype.card n - 1) * det (updateRow M j u) := MultilinearMap.map_update_smul_left _ M j s u @[deprecated (since := "2024-11-03")] alias det_updateRow_smul' := det_updateRow_smul_left -theorem det_updateColumn_smul_left (M : Matrix n n R) (j : n) (s : R) (u : n → R) : - det (updateColumn (s • M) j u) = s ^ (Fintype.card n - 1) * det (updateColumn M j u) := by +theorem det_updateCol_smul_left (M : Matrix n n R) (j : n) (s : R) (u : n → R) : + det (updateCol (s • M) j u) = s ^ (Fintype.card n - 1) * det (updateCol M j u) := by rw [← det_transpose, ← updateRow_transpose, transpose_smul, det_updateRow_smul_left] simp [updateRow_transpose, det_transpose] -@[deprecated (since := "2024-11-03")] alias det_updateColumn_smul' := det_updateColumn_smul_left +@[deprecated (since := "2024-12-11")] alias det_updateColumn_smul' := det_updateCol_smul_left +@[deprecated (since := "2024-12-11")] alias det_updateColumn_smul_left := det_updateCol_smul_left theorem det_updateRow_sum_aux (M : Matrix n n R) {j : n} (s : Finset n) (hj : j ∉ s) (c : n → R) (a : R) : @@ -405,12 +412,14 @@ theorem det_updateRow_sum (A : Matrix n n R) (j : n) (c : n → R) : /-- If we replace a column of a matrix by a linear combination of its columns, then the determinant is multiplied by the coefficient of that column. -/ -theorem det_updateColumn_sum (A : Matrix n n R) (j : n) (c : n → R) : - (A.updateColumn j (fun k ↦ ∑ i, (c i) • A k i)).det = (c j) • A.det := by +theorem det_updateCol_sum (A : Matrix n n R) (j : n) (c : n → R) : + (A.updateCol j (fun k ↦ ∑ i, (c i) • A k i)).det = (c j) • A.det := by rw [← det_transpose, ← updateRow_transpose, ← det_transpose A] convert det_updateRow_sum A.transpose j c simp only [smul_eq_mul, Finset.sum_apply, Pi.smul_apply, transpose_apply] +@[deprecated (since := "2024-12-11")] alias det_updateColumn_sum := det_updateCol_sum + section DetEq /-! ### `det_eq` section @@ -438,21 +447,26 @@ theorem det_updateRow_add_self (A : Matrix n n R) {i j : n} (hij : i ≠ j) : simp [det_updateRow_add, det_zero_of_row_eq hij (updateRow_self.trans (updateRow_ne hij.symm).symm)] -theorem det_updateColumn_add_self (A : Matrix n n R) {i j : n} (hij : i ≠ j) : - det (updateColumn A i fun k => A k i + A k j) = det A := by +theorem det_updateCol_add_self (A : Matrix n n R) {i j : n} (hij : i ≠ j) : + det (updateCol A i fun k => A k i + A k j) = det A := by rw [← det_transpose, ← updateRow_transpose, ← det_transpose A] exact det_updateRow_add_self Aᵀ hij +@[deprecated (since := "2024-12-11")] alias det_updateColumn_add_self := det_updateCol_add_self + theorem det_updateRow_add_smul_self (A : Matrix n n R) {i j : n} (hij : i ≠ j) (c : R) : det (updateRow A i (A i + c • A j)) = det A := by simp [det_updateRow_add, det_updateRow_smul, det_zero_of_row_eq hij (updateRow_self.trans (updateRow_ne hij.symm).symm)] -theorem det_updateColumn_add_smul_self (A : Matrix n n R) {i j : n} (hij : i ≠ j) (c : R) : - det (updateColumn A i fun k => A k i + c • A k j) = det A := by +theorem det_updateCol_add_smul_self (A : Matrix n n R) {i j : n} (hij : i ≠ j) (c : R) : + det (updateCol A i fun k => A k i + c • A k j) = det A := by rw [← det_transpose, ← updateRow_transpose, ← det_transpose A] exact det_updateRow_add_smul_self Aᵀ hij c +@[deprecated (since := "2024-12-11")] +alias det_updateColumn_add_smul_self := det_updateCol_add_smul_self + theorem det_eq_of_forall_row_eq_smul_add_const_aux {A B : Matrix n n R} {s : Finset n} : ∀ (c : n → R) (_ : ∀ i, i ∉ s → c i = 0) (k : n) (_ : k ∉ s) (_ : ∀ i j, A i j = B i j + c i * B k j), det A = det B := by diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean index 8300120633c79..712f006504c42 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean @@ -158,20 +158,26 @@ lemma one_fromRows_isTotallyUnimodular_iff [DecidableEq n] (A : Matrix m n R) : aesop rw [hA, reindex_isTotallyUnimodular, fromRows_one_isTotallyUnimodular_iff] -lemma fromColumns_one_isTotallyUnimodular_iff [DecidableEq m] (A : Matrix m n R) : - (fromColumns A (1 : Matrix m m R)).IsTotallyUnimodular ↔ A.IsTotallyUnimodular := by - rw [←transpose_isTotallyUnimodular_iff, transpose_fromColumns, transpose_one, +lemma fromCols_one_isTotallyUnimodular_iff [DecidableEq m] (A : Matrix m n R) : + (fromCols A (1 : Matrix m m R)).IsTotallyUnimodular ↔ A.IsTotallyUnimodular := by + rw [←transpose_isTotallyUnimodular_iff, transpose_fromCols, transpose_one, fromRows_one_isTotallyUnimodular_iff, transpose_isTotallyUnimodular_iff] -lemma one_fromColumns_isTotallyUnimodular_iff [DecidableEq m] (A : Matrix m n R) : - (fromColumns (1 : Matrix m m R) A).IsTotallyUnimodular ↔ A.IsTotallyUnimodular := by - rw [←transpose_isTotallyUnimodular_iff, transpose_fromColumns, transpose_one, +@[deprecated (since := "2024-12-11")] +alias fromColumns_one_isTotallyUnimodular_iff := fromCols_one_isTotallyUnimodular_iff + +lemma one_fromCols_isTotallyUnimodular_iff [DecidableEq m] (A : Matrix m n R) : + (fromCols (1 : Matrix m m R) A).IsTotallyUnimodular ↔ A.IsTotallyUnimodular := by + rw [←transpose_isTotallyUnimodular_iff, transpose_fromCols, transpose_one, one_fromRows_isTotallyUnimodular_iff, transpose_isTotallyUnimodular_iff] +@[deprecated (since := "2024-12-11")] +alias one_fromColumns_isTotallyUnimodular_iff := one_fromCols_isTotallyUnimodular_iff + alias ⟨_, IsTotallyUnimodular.fromRows_one⟩ := fromRows_one_isTotallyUnimodular_iff alias ⟨_, IsTotallyUnimodular.one_fromRows⟩ := one_fromRows_isTotallyUnimodular_iff -alias ⟨_, IsTotallyUnimodular.fromColumns_one⟩ := fromColumns_one_isTotallyUnimodular_iff -alias ⟨_, IsTotallyUnimodular.one_fromColumns⟩ := one_fromColumns_isTotallyUnimodular_iff +alias ⟨_, IsTotallyUnimodular.fromCols_one⟩ := fromCols_one_isTotallyUnimodular_iff +alias ⟨_, IsTotallyUnimodular.one_fromCols⟩ := one_fromCols_isTotallyUnimodular_iff lemma fromRows_row0_isTotallyUnimodular_iff (A : Matrix m n R) : (fromRows A (row m' 0)).IsTotallyUnimodular ↔ A.IsTotallyUnimodular := by @@ -182,9 +188,12 @@ lemma fromRows_row0_isTotallyUnimodular_iff (A : Matrix m n R) : ext x simp [Pi.single_apply] -lemma fromColumns_col0_isTotallyUnimodular_iff (A : Matrix m n R) : - (fromColumns A (col n' 0)).IsTotallyUnimodular ↔ A.IsTotallyUnimodular := by - rw [← transpose_isTotallyUnimodular_iff, transpose_fromColumns, transpose_col, +lemma fromCols_col0_isTotallyUnimodular_iff (A : Matrix m n R) : + (fromCols A (col n' 0)).IsTotallyUnimodular ↔ A.IsTotallyUnimodular := by + rw [← transpose_isTotallyUnimodular_iff, transpose_fromCols, transpose_col, fromRows_row0_isTotallyUnimodular_iff, transpose_isTotallyUnimodular_iff] +@[deprecated (since := "2024-12-11")] +alias fromColumns_col0_isTotallyUnimodular_iff := fromCols_col0_isTotallyUnimodular_iff + end Matrix diff --git a/Mathlib/LinearAlgebra/Matrix/Permanent.lean b/Mathlib/LinearAlgebra/Matrix/Permanent.lean index 59b4a3a2b020d..38b6d0b855c68 100644 --- a/Mathlib/LinearAlgebra/Matrix/Permanent.lean +++ b/Mathlib/LinearAlgebra/Matrix/Permanent.lean @@ -93,18 +93,21 @@ theorem permanent_smul (M : Matrix n n R) (c : R) : exact prod_mul_pow_card.symm @[simp] -theorem permanent_updateColumn_smul (M : Matrix n n R) (j : n) (c : R) (u : n → R) : - permanent (updateColumn M j <| c • u) = c * permanent (updateColumn M j u) := by - simp only [permanent, ← mul_prod_erase _ _ (mem_univ j), updateColumn_self, Pi.smul_apply, +theorem permanent_updateCol_smul (M : Matrix n n R) (j : n) (c : R) (u : n → R) : + permanent (updateCol M j <| c • u) = c * permanent (updateCol M j u) := by + simp only [permanent, ← mul_prod_erase _ _ (mem_univ j), updateCol_self, Pi.smul_apply, smul_eq_mul, mul_sum, ← mul_assoc] congr 1 with p rw [Finset.prod_congr rfl (fun i hi ↦ ?_)] - simp only [ne_eq, ne_of_mem_erase hi, not_false_eq_true, updateColumn_ne] + simp only [ne_eq, ne_of_mem_erase hi, not_false_eq_true, updateCol_ne] + +@[deprecated (since := "2024-12-11")] +alias permanent_updateColumn_smul := permanent_updateCol_smul @[simp] theorem permanent_updateRow_smul (M : Matrix n n R) (j : n) (c : R) (u : n → R) : permanent (updateRow M j <| c • u) = c * permanent (updateRow M j u) := by - rw [← permanent_transpose, ← updateColumn_transpose, permanent_updateColumn_smul, - updateColumn_transpose, permanent_transpose] + rw [← permanent_transpose, ← updateCol_transpose, permanent_updateCol_smul, + updateCol_transpose, permanent_transpose] end Matrix diff --git a/Mathlib/RingTheory/DedekindDomain/Different.lean b/Mathlib/RingTheory/DedekindDomain/Different.lean index bbc06fe1ddaf3..be9c17f20d9c5 100644 --- a/Mathlib/RingTheory/DedekindDomain/Different.lean +++ b/Mathlib/RingTheory/DedekindDomain/Different.lean @@ -188,7 +188,7 @@ lemma isIntegral_discr_mul_of_mem_traceDual rw [cramer_apply] apply IsIntegral.det intros j k - rw [updateColumn_apply] + rw [updateCol_apply] split · rw [mul_assoc] rw [mem_traceDual_iff_isIntegral] at hx diff --git a/Mathlib/RingTheory/Discriminant.lean b/Mathlib/RingTheory/Discriminant.lean index f3faffa03b668..cd0db7dd94243 100644 --- a/Mathlib/RingTheory/Discriminant.lean +++ b/Mathlib/RingTheory/Discriminant.lean @@ -288,9 +288,9 @@ theorem discr_mul_isIntegral_mem_adjoin [Algebra.IsSeparable K L] [IsIntegrallyC refine Subalgebra.sum_mem _ fun σ _ => Subalgebra.zsmul_mem _ (Subalgebra.prod_mem _ fun j _ => ?_) _ by_cases hji : j = i - · simp only [updateColumn_apply, hji, eq_self_iff_true, PowerBasis.coe_basis] + · simp only [updateCol_apply, hji, eq_self_iff_true, PowerBasis.coe_basis] exact mem_bot.2 (IsIntegrallyClosed.isIntegral_iff.1 <| isIntegral_trace (hz.mul <| hint.pow _)) - · simp only [updateColumn_apply, hji, PowerBasis.coe_basis] + · simp only [updateCol_apply, hji, PowerBasis.coe_basis] exact mem_bot.2 (IsIntegrallyClosed.isIntegral_iff.1 <| isIntegral_trace <| (hint.pow _).mul (hint.pow _)) diff --git a/Mathlib/Topology/Instances/Matrix.lean b/Mathlib/Topology/Instances/Matrix.lean index dc4f941c6159d..6dbd675168fb4 100644 --- a/Mathlib/Topology/Instances/Matrix.lean +++ b/Mathlib/Topology/Instances/Matrix.lean @@ -179,13 +179,16 @@ theorem Continuous.matrix_det [Fintype n] [DecidableEq n] [CommRing R] [Topologi exact continuous_finset_prod _ fun l _ => hA.matrix_elem _ _ @[continuity] -theorem Continuous.matrix_updateColumn [DecidableEq n] (i : n) {A : X → Matrix m n R} +theorem Continuous.matrix_updateCol [DecidableEq n] (i : n) {A : X → Matrix m n R} {B : X → m → R} (hA : Continuous A) (hB : Continuous B) : - Continuous fun x => (A x).updateColumn i (B x) := + Continuous fun x => (A x).updateCol i (B x) := continuous_matrix fun _j k => (continuous_apply k).comp <| ((continuous_apply _).comp hA).update i ((continuous_apply _).comp hB) +@[deprecated (since := "2024-12-11")] +alias Continuous.matrix_updateColumn := Continuous.matrix_updateCol + @[continuity] theorem Continuous.matrix_updateRow [DecidableEq m] (i : m) {A : X → Matrix m n R} {B : X → n → R} (hA : Continuous A) (hB : Continuous B) : Continuous fun x => (A x).updateRow i (B x) := @@ -195,13 +198,13 @@ theorem Continuous.matrix_updateRow [DecidableEq m] (i : m) {A : X → Matrix m theorem Continuous.matrix_cramer [Fintype n] [DecidableEq n] [CommRing R] [TopologicalRing R] {A : X → Matrix n n R} {B : X → n → R} (hA : Continuous A) (hB : Continuous B) : Continuous fun x => cramer (A x) (B x) := - continuous_pi fun _ => (hA.matrix_updateColumn _ hB).matrix_det + continuous_pi fun _ => (hA.matrix_updateCol _ hB).matrix_det @[continuity] theorem Continuous.matrix_adjugate [Fintype n] [DecidableEq n] [CommRing R] [TopologicalRing R] {A : X → Matrix n n R} (hA : Continuous A) : Continuous fun x => (A x).adjugate := continuous_matrix fun _j k => - (hA.matrix_transpose.matrix_updateColumn k continuous_const).matrix_det + (hA.matrix_transpose.matrix_updateCol k continuous_const).matrix_det /-- When `Ring.inverse` is continuous at the determinant (such as in a `NormedRing`, or a topological field), so is `Matrix.inv`. -/ diff --git a/scripts/nolints_prime_decls.txt b/scripts/nolints_prime_decls.txt index 414151f344621..141e187da20d1 100644 --- a/scripts/nolints_prime_decls.txt +++ b/scripts/nolints_prime_decls.txt @@ -2563,7 +2563,7 @@ Matrix.cons_val_succ' Matrix.cons_val_zero' Matrix.det_apply' Matrix.det_units_conj' -Matrix.det_updateColumn_smul' +Matrix.det_updateCol_smul' Matrix.det_updateRow_smul' Matrix.diagonal_apply_ne' Matrix.diagonal_intCast' From f7bc890fd36380b7c39fe2fbe87836fd44cde175 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Wed, 11 Dec 2024 08:26:30 +0000 Subject: [PATCH 656/829] feat(Topology/ContinuousMap): "ContinuousMap.evalCLM" without compactness (#19876) `ContinuousMap.evalCLM` is the evaluation-at-a-point functional on continuous-function spaces, a useful gadget to have. This is currently only defined when the domain is compact (presumably for historical reasons). This PR moves it around slightly so it is now available for any domain space. Co-authored-by: Eric Wieser --- Mathlib/Topology/ContinuousMap/Algebra.lean | 7 +++++++ Mathlib/Topology/ContinuousMap/Bounded/Basic.lean | 4 +--- Mathlib/Topology/ContinuousMap/Compact.lean | 8 -------- Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean | 10 +++++----- 4 files changed, 13 insertions(+), 16 deletions(-) diff --git a/Mathlib/Topology/ContinuousMap/Algebra.lean b/Mathlib/Topology/ContinuousMap/Algebra.lean index 095d12e987b09..58c98d954d09d 100644 --- a/Mathlib/Topology/ContinuousMap/Algebra.lean +++ b/Mathlib/Topology/ContinuousMap/Algebra.lean @@ -617,6 +617,13 @@ def coeFnLinearMap : C(α, M) →ₗ[R] α → M := { (coeFnAddMonoidHom : C(α, M) →+ _) with map_smul' := coe_smul } +/-- Evaluation at a point, as a continuous linear map. -/ +@[simps apply] +def evalCLM (x : α) : C(α, M) →L[R] M where + toFun f := f x + map_add' _ _ := add_apply _ _ x + map_smul' _ _ := smul_apply _ _ x + end ContinuousMap end ModuleStructure diff --git a/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean b/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean index 5c2ccedeed4ea..b378f1cd1420d 100644 --- a/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean +++ b/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean @@ -1087,14 +1087,12 @@ instance instModule : Module 𝕜 (α →ᵇ β) := variable (𝕜) /-- The evaluation at a point, as a continuous linear map from `α →ᵇ β` to `β`. -/ +@[simps] def evalCLM (x : α) : (α →ᵇ β) →L[𝕜] β where toFun f := f x map_add' _ _ := add_apply _ _ map_smul' _ _ := smul_apply _ _ _ -@[simp] -theorem evalCLM_apply (x : α) (f : α →ᵇ β) : evalCLM 𝕜 x f = f x := rfl - variable (α β) /-- The linear map forgetting that a bounded continuous function is bounded. -/ diff --git a/Mathlib/Topology/ContinuousMap/Compact.lean b/Mathlib/Topology/ContinuousMap/Compact.lean index e6cb2f048e863..1e69d46a45e39 100644 --- a/Mathlib/Topology/ContinuousMap/Compact.lean +++ b/Mathlib/Topology/ContinuousMap/Compact.lean @@ -301,14 +301,6 @@ def linearIsometryBoundedOfCompact : C(α, E) ≃ₗᵢ[𝕜] α →ᵇ E := norm_cast norm_map' := fun _ => rfl } -variable {α E} - --- to match `BoundedContinuousFunction.evalCLM` -/-- The evaluation at a point, as a continuous linear map from `C(α, 𝕜)` to `𝕜`. -/ -def evalCLM (x : α) : C(α, E) →L[𝕜] E := - (BoundedContinuousFunction.evalCLM 𝕜 x).comp - (linearIsometryBoundedOfCompact α E 𝕜).toLinearIsometry.toContinuousLinearMap - end -- this lemma and the next are the analogues of those autogenerated by `@[simps]` for diff --git a/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean b/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean index e69cf73093b84..d354b286c82fb 100644 --- a/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean +++ b/Mathlib/Topology/ContinuousMap/ContinuousMapZero.lean @@ -230,13 +230,13 @@ def toContinuousMapCLM (M : Type*) [Semiring M] [Module M R] [ContinuousConstSMu map_smul' _ _ := rfl /-- The evaluation at a point, as a continuous linear map from `C(X, R)₀` to `R`. -/ -def evalCLM (𝕜 : Type*) {R : Type*} [CompactSpace X] [NormedField 𝕜] [NormedCommRing R] - [NormedSpace 𝕜 R] (x : X) : C(X, R)₀ →L[𝕜] R := - (ContinuousMap.evalCLM 𝕜 x).comp (toContinuousMapCLM 𝕜 : C(X, R)₀ →L[𝕜] C(X, R)) +def evalCLM (𝕜 : Type*) [Semiring 𝕜] [Module 𝕜 R] [ContinuousConstSMul 𝕜 R] (x : X) : + C(X, R)₀ →L[𝕜] R := + (ContinuousMap.evalCLM 𝕜 x).comp (toContinuousMapCLM 𝕜) @[simp] -lemma evalCLM_apply {𝕜 : Type*} {R : Type*} [CompactSpace X] [NormedField 𝕜] [NormedCommRing R] - [NormedSpace 𝕜 R] (x : X) (f : C(X, R)₀) : evalCLM 𝕜 x f = f x := rfl +lemma evalCLM_apply {𝕜 : Type*} [Semiring 𝕜] [Module 𝕜 R] [ContinuousConstSMul 𝕜 R] + (x : X) (f : C(X, R)₀) : evalCLM 𝕜 x f = f x := rfl /-- Coercion to a function as an `AddMonoidHom`. Similar to `ContinuousMap.coeFnAddMonoidHom`. -/ def coeFnAddMonoidHom : C(X, R)₀ →+ X → R where From d840f50ab07fb0170966e100a11aef6123eaf306 Mon Sep 17 00:00:00 2001 From: David Renshaw Date: Wed, 11 Dec 2024 08:26:31 +0000 Subject: [PATCH 657/829] chore(NumberTheory/ModularForms/JacobiTheta): simplify proof of isBigO_atTop_F_int_zero_sub (#19879) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * `have ha' : (a : UnitAddCircle) = 0 ↔ a = 0` can be directly proved by `AddCircle.coe_eq_zero_iff_of_mem_Ico ha`. * That makes the proof small enough that it makes sense to inline it at its use in `simp_rw`. This simplification was found by [`tryAtEachStep`](https://github.com/dwrensha/tryAtEachStep). --- Mathlib/NumberTheory/ModularForms/JacobiTheta/Bounds.lean | 6 +----- 1 file changed, 1 insertion(+), 5 deletions(-) diff --git a/Mathlib/NumberTheory/ModularForms/JacobiTheta/Bounds.lean b/Mathlib/NumberTheory/ModularForms/JacobiTheta/Bounds.lean index 4e6f7d202b007..b106e2e0d0de9 100644 --- a/Mathlib/NumberTheory/ModularForms/JacobiTheta/Bounds.lean +++ b/Mathlib/NumberTheory/ModularForms/JacobiTheta/Bounds.lean @@ -249,11 +249,7 @@ lemma isBigO_atTop_F_int_zero_sub (a : UnitAddCircle) : ∃ p, 0 < p ∧ obtain ⟨a, ha, rfl⟩ := a.eq_coe_Ico obtain ⟨p, hp, hp'⟩ := isBigO_atTop_F_nat_zero_sub ha.1 obtain ⟨q, hq, hq'⟩ := isBigO_atTop_F_nat_zero_sub (sub_nonneg.mpr ha.2.le) - have ha' : (a : UnitAddCircle) = 0 ↔ a = 0 := by - rw [← AddCircle.coe_eq_coe_iff_of_mem_Ico (hp := ⟨zero_lt_one' ℝ⟩), QuotientAddGroup.mk_zero] - · rw [zero_add]; exact ha - · simp - simp_rw [ha'] + simp_rw [AddCircle.coe_eq_zero_iff_of_mem_Ico ha] simp_rw [eq_false_intro (by linarith [ha.2] : 1 - a ≠ 0), if_false, sub_zero] at hq' refine ⟨_, lt_min hp hq, ?_⟩ have : (fun t ↦ F_int 0 a t - (if a = 0 then 1 else 0)) =ᶠ[atTop] From 868f3499d72c038be8f2e1af09339d67204e6153 Mon Sep 17 00:00:00 2001 From: David Renshaw Date: Wed, 11 Dec 2024 09:46:55 +0000 Subject: [PATCH 658/829] chore(Analysis/SpecialFunctions/Integrals): simplify proof of intervalIntegrable_cpow (#19877) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * `have : Ioc c 0 = Ioo c 0 ∪ {(0 : ℝ)} ` is directly proved by `(Ioo_union_right hc).symm` * That shorter proof can then be inlined at its use in the `simp` in the next line. This simplification was found by [`tryAtEachStep`](https://github.com/dwrensha/tryAtEachStep). --- Mathlib/Analysis/SpecialFunctions/Integrals.lean | 5 +---- 1 file changed, 1 insertion(+), 4 deletions(-) diff --git a/Mathlib/Analysis/SpecialFunctions/Integrals.lean b/Mathlib/Analysis/SpecialFunctions/Integrals.lean index f80d1afa00a62..b3b78b374fb6f 100644 --- a/Mathlib/Analysis/SpecialFunctions/Integrals.lean +++ b/Mathlib/Analysis/SpecialFunctions/Integrals.lean @@ -140,10 +140,7 @@ theorem intervalIntegrable_cpow {r : ℂ} (h : 0 ≤ r.re ∨ (0 : ℝ) ∉ [[a, · -- case `c < 0`: integrand is identically constant, *except* at `x = 0` if `r ≠ 0`. apply IntervalIntegrable.symm rw [intervalIntegrable_iff_integrableOn_Ioc_of_le hc.le] - have : Ioc c 0 = Ioo c 0 ∪ {(0 : ℝ)} := by - rw [← Ioo_union_Icc_eq_Ioc hc (le_refl 0), ← Icc_def] - simp_rw [← le_antisymm_iff, setOf_eq_eq_singleton'] - rw [this, integrableOn_union, and_comm]; constructor + rw [← Ioo_union_right hc, integrableOn_union, and_comm]; constructor · refine integrableOn_singleton_iff.mpr (Or.inr ?_) exact isFiniteMeasureOnCompacts_of_isLocallyFiniteMeasure.lt_top_of_isCompact isCompact_singleton From 69471a419f2182e293a215c7c45a25218f345518 Mon Sep 17 00:00:00 2001 From: David Renshaw Date: Wed, 11 Dec 2024 09:57:13 +0000 Subject: [PATCH 659/829] chore(Algebra/Lie/Quotient): remove unused variable 'I' (#19864) Removes an unused `variable {I}` declaration which causes the `I` to sometimes show up as a hypothesis whose type is a metavariable. When this file was originally ported in #4822, the variable `I` was declared to have type `LieIdeal R L` and then the lower-down `variable {N I}` declaration served to modify the implicitness of that existing variable: ```lean ... variable (N N' : LieSubmodule R L M) (I J : LieIdeal R L) ... namespace Quotient variable {N I} ... ``` However, in #15538, the initial (typed) declaration of `I` was removed. This had the effect of causing an `I` to appear with a metavarable in some contexts, e.g. the proof of `lieQuotientLieRing.leibniz_lie`: ```lean I : ?m.48970 ``` I noticed this problem when I tried running `exact?` in the proof of `lieQuotientLieRing.leibniz_lie`; doing so returned a nonsensical result. --- Mathlib/Algebra/Lie/Quotient.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Lie/Quotient.lean b/Mathlib/Algebra/Lie/Quotient.lean index e6ae53a34c3e9..367d16adeac24 100644 --- a/Mathlib/Algebra/Lie/Quotient.lean +++ b/Mathlib/Algebra/Lie/Quotient.lean @@ -43,7 +43,7 @@ instance : HasQuotient M (LieSubmodule R L M) := namespace Quotient -variable {N I} +variable {N} instance addCommGroup : AddCommGroup (M ⧸ N) := Submodule.Quotient.addCommGroup _ From df8b92afb7f8ff7c452038478fddaab841f07d8b Mon Sep 17 00:00:00 2001 From: Lorenzo Luccioli Date: Wed, 11 Dec 2024 13:07:18 +0000 Subject: [PATCH 660/829] feat: measurability of addition and multiplication on `EReal` (#17097) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add a section about `EReal` in the file `MeasureTheory.Constructions.BorelSpace.Real.lean` with the instances for the measurability of the sum and multiplication between two extended real numbers as well as some lemmas that are needed for the proofs. Co-authored-by: Pietro Monticone <38562595+pitmonticone@users.noreply.github.com> Co-authored-by: Yaël Dillies --- .../Constructions/BorelSpace/Real.lean | 67 +++++++++++++++++++ Mathlib/Topology/Instances/EReal.lean | 9 +++ 2 files changed, 76 insertions(+) diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean index 0d80172e45051..5ed34323264cf 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Yury Kudryashov -/ import Mathlib.MeasureTheory.Constructions.BorelSpace.Order +import Mathlib.MeasureTheory.MeasurableSpace.Prod /-! # Borel (measurable) spaces ℝ, ℝ≥0, ℝ≥0∞ @@ -465,6 +466,72 @@ _root_.measurable_of_tendsto_nnreal := NNReal.measurable_of_tendsto end NNReal +namespace EReal + +lemma measurableEmbedding_coe : MeasurableEmbedding Real.toEReal := + isOpenEmbedding_coe.measurableEmbedding + +instance : MeasurableAdd₂ EReal := ⟨EReal.lowerSemicontinuous_add.measurable⟩ + +section MeasurableMul + +variable {α β γ : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} {mγ : MeasurableSpace γ} + +lemma measurable_of_real_prod {f : EReal × β → γ} + (h_real : Measurable fun p : ℝ × β ↦ f (p.1, p.2)) + (h_bot : Measurable fun x ↦ f (⊥, x)) (h_top : Measurable fun x ↦ f (⊤, x)) : Measurable f := + .of_union₃_range_cover (measurableEmbedding_prod_mk_left _) (measurableEmbedding_prod_mk_left _) + (measurableEmbedding_coe.prod_mk .id) (by simp [-univ_subset_iff, subset_def, EReal.forall]) + h_bot h_top h_real + +lemma measurable_of_real_real {f : EReal × EReal → β} + (h_real : Measurable fun p : ℝ × ℝ ↦ f (p.1, p.2)) + (h_bot_left : Measurable fun r : ℝ ↦ f (⊥, r)) + (h_top_left : Measurable fun r : ℝ ↦ f (⊤, r)) + (h_bot_right : Measurable fun r : ℝ ↦ f (r, ⊥)) + (h_top_right : Measurable fun r : ℝ ↦ f (r, ⊤)) : + Measurable f := by + refine measurable_of_real_prod ?_ ?_ ?_ + · refine measurable_swap_iff.mp <| measurable_of_real_prod ?_ h_bot_right h_top_right + exact h_real.comp measurable_swap + · exact measurable_of_measurable_real h_bot_left + · exact measurable_of_measurable_real h_top_left + +private lemma measurable_const_mul (c : EReal) : Measurable fun (x : EReal) ↦ c * x := by + refine measurable_of_measurable_real ?_ + have h1 : (fun (p : ℝ) ↦ (⊥ : EReal) * p) + = fun p ↦ if p = 0 then (0 : EReal) else (if p < 0 then ⊤ else ⊥) := by + ext p + split_ifs with h1 h2 + · simp [h1] + · rw [bot_mul_coe_of_neg h2] + · rw [bot_mul_coe_of_pos] + exact lt_of_le_of_ne (not_lt.mp h2) (Ne.symm h1) + have h2 : Measurable fun (p : ℝ) ↦ if p = 0 then (0 : EReal) else if p < 0 then ⊤ else ⊥ := by + refine Measurable.piecewise (measurableSet_singleton _) measurable_const ?_ + exact Measurable.piecewise measurableSet_Iio measurable_const measurable_const + induction c with + | h_bot => rwa [h1] + | h_real c => exact (measurable_id.const_mul _).coe_real_ereal + | h_top => + simp_rw [← neg_bot, neg_mul] + apply Measurable.neg + rwa [h1] + +instance : MeasurableMul₂ EReal := by + refine ⟨measurable_of_real_real ?_ ?_ ?_ ?_ ?_⟩ + · exact (measurable_fst.mul measurable_snd).coe_real_ereal + · exact (measurable_const_mul _).comp measurable_coe_real_ereal + · exact (measurable_const_mul _).comp measurable_coe_real_ereal + · simp_rw [mul_comm _ ⊥] + exact (measurable_const_mul _).comp measurable_coe_real_ereal + · simp_rw [mul_comm _ ⊤] + exact (measurable_const_mul _).comp measurable_coe_real_ereal + +end MeasurableMul + +end EReal + /-- If a function `f : α → ℝ≥0` is measurable and the measure is σ-finite, then there exists spanning measurable sets with finite measure on which `f` is bounded. See also `StronglyMeasurable.exists_spanning_measurableSet_norm_le` for functions into normed diff --git a/Mathlib/Topology/Instances/EReal.lean b/Mathlib/Topology/Instances/EReal.lean index 746c9fd365e5b..3af0b2ab4b18d 100644 --- a/Mathlib/Topology/Instances/EReal.lean +++ b/Mathlib/Topology/Instances/EReal.lean @@ -7,6 +7,7 @@ import Mathlib.Data.Rat.Encodable import Mathlib.Data.Real.EReal import Mathlib.Topology.Instances.ENNReal import Mathlib.Topology.Order.MonotoneContinuity +import Mathlib.Topology.Semicontinuous /-! # Topological structure on `EReal` @@ -441,4 +442,12 @@ theorem continuousAt_mul {p : EReal × EReal} (h₁ : p.1 ≠ 0 ∨ p.2 ≠ ⊥) exact continuousAt_mul_top_ne_zero h₄ · exact continuousAt_mul_top_top +lemma lowerSemicontinuous_add : LowerSemicontinuous fun p : EReal × EReal ↦ p.1 + p.2 := by + intro x y + by_cases hx₁ : x.1 = ⊥ + · simp [hx₁] + by_cases hx₂ : x.2 = ⊥ + · simp [hx₂] + · exact continuousAt_add (.inr hx₂) (.inl hx₁) |>.lowerSemicontinuousAt _ + end EReal From a1c52d244b6ba4b5b6e16475841e9946cb9ece13 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 11 Dec 2024 13:42:49 +0000 Subject: [PATCH 661/829] chore(discover-lean-pr-testing): Zulip post with command for merging lean-testing PRs (#19857) Also make the script for merging the lean-testing PRs a bit more robust: write sed output to a new file, and move that to the original, instead of doing inplace updates with `sed -i` (which doesn't seem to work wel on MacOS). --- .../workflows/discover-lean-pr-testing.yml | 22 ++++++++++++++----- scripts/merge-lean-testing-pr.sh | 3 ++- 2 files changed, 19 insertions(+), 6 deletions(-) diff --git a/.github/workflows/discover-lean-pr-testing.yml b/.github/workflows/discover-lean-pr-testing.yml index 04b70247d680f..0abbbef132b36 100644 --- a/.github/workflows/discover-lean-pr-testing.yml +++ b/.github/workflows/discover-lean-pr-testing.yml @@ -124,6 +124,22 @@ jobs: echo "branches_exist=true" >> "$GITHUB_ENV" fi + - name: Prepare Zulip message + id: zulip-message + run: | + BRANCHES="${{ steps.find-branches.outputs.branches }}" + { + printf "msg<> "$GITHUB_OUTPUT" + - name: Send message on Zulip if: env.branches_exist == 'true' uses: zulip/github-actions-zulip/send-message@v1 @@ -135,8 +151,4 @@ jobs: type: 'stream' topic: 'Mergeable lean testing PRs' content: | - We will need to merge the following branches into `nightly-testing`: - - ${{ steps.find-branches.outputs.branches }} - - This can be done using the script `scripts/merge-lean-testing-pr.sh`, which takes 1 lean-PR number as argument. + ${{ steps.zulip-message.outputs.msg }} diff --git a/scripts/merge-lean-testing-pr.sh b/scripts/merge-lean-testing-pr.sh index 7a07419dbe47f..9cd3685a83a41 100755 --- a/scripts/merge-lean-testing-pr.sh +++ b/scripts/merge-lean-testing-pr.sh @@ -17,7 +17,8 @@ if ! git merge origin/$BRANCH_NAME; then git add lean-toolchain lakefile.lean lake-manifest.json fi -sed -i "s/$BRANCH_NAME/nightly-testing/g" lakefile.lean +sed "s/$BRANCH_NAME/nightly-testing/g" < lakefile.lean > lakefile.lean.new +mv lakefile.lean.new lakefile.lean git add lakefile.lean # Check for merge conflicts From d33492867a63c85a24d0f0a82a5e98bfaf4f2e0c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Wed, 11 Dec 2024 13:51:00 +0000 Subject: [PATCH 662/829] feat(CategoryTheory): construction of sections of functors by transfinite induction (#19330) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Given a functor `F : Jᵒᵖ ⥤ Type v` where `J` is a well-ordered type, we introduce a structure `F.WellOrderInductionData` which allows to show that the map `F.sections → F.obj (op ⊥)` is surjective. The data and properties in `F.WellOrderInductionData` consist of a section to the maps `F.obj (op (Order.succ j)) → F.obj (op j)` when `j` is not maximal, and, when `j` is limit, a section to the canonical map from `F.obj (op j)` to the type of compatible families of elements in `F.obj (op i)` for `i < j`. In other words, from `val₀ : F.obj (op ⊥)`, a term `d : F.WellOrderInductionData` allows the construction, by transfinite induction, of a section of `F` which restricts to `val₀`. --- Mathlib.lean | 1 + .../SmallObject/WellOrderInductionData.lean | 272 ++++++++++++++++++ 2 files changed, 273 insertions(+) create mode 100644 Mathlib/CategoryTheory/SmallObject/WellOrderInductionData.lean diff --git a/Mathlib.lean b/Mathlib.lean index 81fbb6c85c38c..8569d32bdc872 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2139,6 +2139,7 @@ import Mathlib.CategoryTheory.SmallObject.Iteration.Basic import Mathlib.CategoryTheory.SmallObject.Iteration.ExtendToSucc import Mathlib.CategoryTheory.SmallObject.Iteration.Nonempty import Mathlib.CategoryTheory.SmallObject.Iteration.UniqueHom +import Mathlib.CategoryTheory.SmallObject.WellOrderInductionData import Mathlib.CategoryTheory.Square import Mathlib.CategoryTheory.Subobject.Basic import Mathlib.CategoryTheory.Subobject.Comma diff --git a/Mathlib/CategoryTheory/SmallObject/WellOrderInductionData.lean b/Mathlib/CategoryTheory/SmallObject/WellOrderInductionData.lean new file mode 100644 index 0000000000000..1c0de001e52da --- /dev/null +++ b/Mathlib/CategoryTheory/SmallObject/WellOrderInductionData.lean @@ -0,0 +1,272 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import Mathlib.CategoryTheory.Category.Preorder +import Mathlib.CategoryTheory.Functor.Category +import Mathlib.CategoryTheory.Types +import Mathlib.Order.SuccPred.Limit + +/-! +# Limits of inverse systems indexed by well-ordered types + +Given a functor `F : Jᵒᵖ ⥤ Type v` where `J` is a well-ordered type, +we introduce a structure `F.WellOrderInductionData` which allows +to show that the map `F.sections → F.obj (op ⊥)` is surjective. + +The data and properties in `F.WellOrderInductionData` consist of a +section to the maps `F.obj (op (Order.succ j)) → F.obj (op j)` when `j` is not maximal, +and, when `j` is limit, a section to the canonical map from `F.obj (op j)` +to the type of compatible families of elements in `F.obj (op i)` for `i < j`. + +In other words, from `val₀ : F.obj (op ⊥)`, a term `d : F.WellOrderInductionData` +allows the construction, by transfinite induction, of a section of `F` +which restricts to `val₀`. + +-/ + +universe v u + +namespace CategoryTheory + +open Opposite + +namespace Functor + +variable {J : Type u} [LinearOrder J] [SuccOrder J] + (F : Jᵒᵖ ⥤ Type v) + +/-- Given a functor `F : Jᵒᵖ ⥤ Type v` where `J` is a well-ordered type, this data +allows to construct a section of `F` from an element in `F.obj (op ⊥)`, +see `WellOrderInductionData.sectionsMk`. -/ +structure WellOrderInductionData where + /-- A section `F.obj (op j) → F.obj (op (Order.succ j))` to the restriction + `F.obj (op (Order.succ j)) → F.obj (op j)` when `j` is not maximal. -/ + succ (j : J) (hj : ¬IsMax j) (x : F.obj (op j)) : F.obj (op (Order.succ j)) + map_succ (j : J) (hj : ¬IsMax j) (x : F.obj (op j)) : + F.map (homOfLE (Order.le_succ j)).op (succ j hj x) = x + /-- When `j` is a limit element, and `x` is a compatible family of elements + in `F.obj (op i)` for all `i < j`, this is a lifting to `F.obj (op j)`. -/ + lift (j : J) (hj : Order.IsSuccLimit j) + (x : ((OrderHom.Subtype.val (· ∈ Set.Iio j)).monotone.functor.op ⋙ F).sections) : + F.obj (op j) + map_lift (j : J) (hj : Order.IsSuccLimit j) + (x : ((OrderHom.Subtype.val (· ∈ Set.Iio j)).monotone.functor.op ⋙ F).sections) + (i : J) (hi : i < j) : + F.map (homOfLE hi.le).op (lift j hj x) = x.val (op ⟨i, hi⟩) + +namespace WellOrderInductionData + +variable {F} (d : F.WellOrderInductionData) [OrderBot J] + +/-- Given `d : F.WellOrderInductionData`, `val₀ : F.obj (op ⊥)` and `j : J`, +this is the data of an element `val : F.obj (op j)` such that the induced +compatible family of elements in all `F.obj (op i)` for `i ≤ j` +is determined by `val₀` and the choice of "liftings" given by `d`. -/ +structure Extension (val₀ : F.obj (op ⊥)) (j : J) where + /-- An element in `F.obj (op j)`, which, by restriction, induces elements + in `F.obj (op i)` for all `i ≤ j`. -/ + val : F.obj (op j) + map_zero : F.map (homOfLE bot_le).op val = val₀ + map_succ (i : J) (hi : i < j) : + F.map (homOfLE (Order.succ_le_of_lt hi)).op val = + d.succ i (not_isMax_iff.2 ⟨_, hi⟩) (F.map (homOfLE hi.le).op val) + map_limit (i : J) (hi : Order.IsSuccLimit i) (hij : i ≤ j) : + F.map (homOfLE hij).op val = d.lift i hi + { val := fun ⟨⟨k, hk⟩⟩ ↦ F.map (homOfLE (hk.le.trans hij)).op val + property := fun f ↦ by + dsimp + rw [← FunctorToTypes.map_comp_apply] + rfl } + +namespace Extension + +variable {d} {val₀ : F.obj (op ⊥)} + +/-- An element in `d.Extension val₀ j` induces an element in `d.Extension val₀ i` when `i ≤ j`. -/ +@[simps] +def ofLE {j : J} (e : d.Extension val₀ j) {i : J} (hij : i ≤ j) : d.Extension val₀ i where + val := F.map (homOfLE hij).op e.val + map_zero := by + rw [← FunctorToTypes.map_comp_apply] + exact e.map_zero + map_succ k hk := by + rw [← FunctorToTypes.map_comp_apply, ← FunctorToTypes.map_comp_apply, ← op_comp, ← op_comp, + homOfLE_comp, homOfLE_comp, e.map_succ k (lt_of_lt_of_le hk hij)] + map_limit k hk hki := by + rw [← FunctorToTypes.map_comp_apply, ← op_comp, homOfLE_comp, + e.map_limit k hk (hki.trans hij)] + congr + ext ⟨l, hl⟩ + dsimp + rw [← FunctorToTypes.map_comp_apply] + rfl + +lemma val_injective {j : J} {e e' : d.Extension val₀ j} (h : e.val = e'.val) : e = e' := by + cases e + cases e' + subst h + rfl + +instance [WellFoundedLT J] (j : J) : Subsingleton (d.Extension val₀ j) := by + induction j using SuccOrder.limitRecOn with + | hm i hi => + obtain rfl : i = ⊥ := by simpa using hi + refine Subsingleton.intro (fun e₁ e₂ ↦ val_injective ?_) + have h₁ := e₁.map_zero + have h₂ := e₂.map_zero + simp only [homOfLE_refl, op_id, FunctorToTypes.map_id_apply] at h₁ h₂ + rw [h₁, h₂] + | hs i hi hi' => + refine Subsingleton.intro (fun e₁ e₂ ↦ val_injective ?_) + have h₁ := e₁.map_succ i (Order.lt_succ_of_not_isMax hi) + have h₂ := e₂.map_succ i (Order.lt_succ_of_not_isMax hi) + simp only [homOfLE_refl, op_id, FunctorToTypes.map_id_apply, homOfLE_leOfHom] at h₁ h₂ + rw [h₁, h₂] + congr + exact congr_arg val + (Subsingleton.elim (e₁.ofLE (Order.le_succ i)) (e₂.ofLE (Order.le_succ i))) + | hl i hi hi' => + refine Subsingleton.intro (fun e₁ e₂ ↦ val_injective ?_) + have h₁ := e₁.map_limit i hi (by rfl) + have h₂ := e₂.map_limit i hi (by rfl) + simp only [homOfLE_refl, op_id, FunctorToTypes.map_id_apply, OrderHom.Subtype.val_coe, + comp_obj, op_obj, Monotone.functor_obj, homOfLE_leOfHom] at h₁ h₂ + rw [h₁, h₂] + congr + ext ⟨⟨l, hl⟩⟩ + have := hi' l hl + exact congr_arg val (Subsingleton.elim (e₁.ofLE hl.le) (e₂.ofLE hl.le)) + +lemma compatibility [WellFoundedLT J] + {j : J} (e : d.Extension val₀ j) {i : J} (e' : d.Extension val₀ i) (h : i ≤ j) : + F.map (homOfLE h).op e.val = e'.val := by + obtain rfl : e' = e.ofLE h := Subsingleton.elim _ _ + rfl + +variable (d val₀) in +/-- The obvious element in `d.Extension val₀ ⊥`. -/ +@[simps] +def zero : d.Extension val₀ ⊥ where + val := val₀ + map_zero := by simp + map_succ i hi := by simp at hi + map_limit i hi hij := by + obtain rfl : i = ⊥ := by simpa using hij + simpa using hi.not_isMin + +/-- The element in `d.Extension val₀ (Order.succ j)` obtained by extending +an element in `d.Extension val₀ j` when `j` is not maximal. -/ +def succ {j : J} (e : d.Extension val₀ j) (hj : ¬IsMax j) : + d.Extension val₀ (Order.succ j) where + val := d.succ j hj e.val + map_zero := by + simp only [← e.map_zero] + conv_rhs => rw [← d.map_succ j hj e.val] + rw [← FunctorToTypes.map_comp_apply] + rfl + map_succ i hi := by + obtain hij | rfl := ((Order.lt_succ_iff_of_not_isMax hj).mp hi).lt_or_eq + · rw [← homOfLE_comp ((Order.lt_succ_iff_of_not_isMax hj).mp hi) (Order.le_succ j), op_comp, + FunctorToTypes.map_comp_apply, d.map_succ, ← e.map_succ i hij, + ← homOfLE_comp (Order.succ_le_of_lt hij) (Order.le_succ j), op_comp, + FunctorToTypes.map_comp_apply, d.map_succ] + · simp only [homOfLE_refl, op_id, FunctorToTypes.map_id_apply, homOfLE_leOfHom, + d.map_succ] + map_limit i hi hij := by + obtain hij | rfl := hij.lt_or_eq + · have hij' : i ≤ j := (Order.lt_succ_iff_of_not_isMax hj).mp hij + have := congr_arg (F.map (homOfLE hij').op) (d.map_succ j hj e.val) + rw [e.map_limit i hi, ← FunctorToTypes.map_comp_apply, ← op_comp, homOfLE_comp] at this + rw [this] + congr + ext ⟨⟨l, hl⟩⟩ + dsimp + conv_lhs => rw [← d.map_succ j hj e.val] + rw [← FunctorToTypes.map_comp_apply] + rfl + · exfalso + exact hj hi.isMax + +variable [WellFoundedLT J] + +/-- When `j` is a limit element, this is the exntesion to `d.Extension val₀ j` +of a family of elements in `d.Extension val₀ i` for all `i < j`. -/ +def limit (j : J) (hj : Order.IsSuccLimit j) + (e : ∀ (i : J) (_ : i < j), d.Extension val₀ i) : + d.Extension val₀ j where + val := d.lift j hj + { val := fun ⟨i, hi⟩ ↦ (e i hi).val + property := fun f ↦ by apply compatibility } + map_zero := by + rw [d.map_lift _ _ _ _ (by simpa [bot_lt_iff_ne_bot] using hj.not_isMin)] + simpa only [homOfLE_refl, op_id, FunctorToTypes.map_id_apply] using + (e ⊥ (by simpa [bot_lt_iff_ne_bot] using hj.not_isMin)).map_zero + map_succ i hi := by + convert (e (Order.succ i) ((Order.IsSuccLimit.succ_lt_iff hj).mpr hi)).map_succ i + (by + simp only [Order.lt_succ_iff_not_isMax, not_isMax_iff] + exact ⟨_, hi⟩) using 1 + · dsimp + rw [FunctorToTypes.map_id_apply, + d.map_lift _ _ _ _ ((Order.IsSuccLimit.succ_lt_iff hj).mpr hi)] + · congr + rw [d.map_lift _ _ _ _ hi] + symm + apply compatibility + map_limit i hi hij := by + obtain hij' | rfl := hij.lt_or_eq + · have := (e i hij').map_limit i hi (by rfl) + dsimp at this ⊢ + rw [FunctorToTypes.map_id_apply] at this + rw [d.map_lift _ _ _ _ hij'] + dsimp + rw [this] + congr + dsimp + ext ⟨⟨l, hl⟩⟩ + rw [map_lift _ _ _ _ _ (hl.trans hij')] + apply compatibility + · dsimp + rw [FunctorToTypes.map_id_apply] + congr + ext ⟨⟨l, hl⟩⟩ + rw [d.map_lift _ _ _ _ hl] + +instance (j : J) : Nonempty (d.Extension val₀ j) := by + induction j using SuccOrder.limitRecOn with + | hm i hi => + obtain rfl : i = ⊥ := by simpa using hi + exact ⟨zero d val₀⟩ + | hs i hi hi' => exact ⟨hi'.some.succ hi⟩ + | hl i hi hi' => exact ⟨limit i hi (fun l hl ↦ (hi' l hl).some)⟩ + +noncomputable instance (j : J) : Unique (d.Extension val₀ j) := + uniqueOfSubsingleton (Nonempty.some inferInstance) + +end Extension + +variable [WellFoundedLT J] + +/-- When `J` is a well-ordered type, `F : Jᵒᵖ ⥤ Type v`, and `d : F.WellOrderInductionData`, +this is the section of `F` that is determined by `val₀ : F.obj (op ⊥)` -/ +noncomputable def sectionsMk (val₀ : F.obj (op ⊥)) : F.sections where + val j := (default : d.Extension val₀ j.unop).val + property := fun f ↦ by apply Extension.compatibility + +lemma sectionsMk_val_op_bot (val₀ : F.obj (op ⊥)) : + (d.sectionsMk val₀).val (op ⊥) = val₀ := by + simpa using (default : d.Extension val₀ ⊥).map_zero + +include d in +lemma surjective : + Function.Surjective ((fun s ↦ s (op ⊥)) ∘ Subtype.val : F.sections → F.obj (op ⊥)) := + fun val₀ ↦ ⟨d.sectionsMk val₀, d.sectionsMk_val_op_bot val₀⟩ + +end WellOrderInductionData + +end Functor + +end CategoryTheory From 6e4ad6d7c04852583809838601c82541222d2082 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Wed, 11 Dec 2024 13:59:18 +0000 Subject: [PATCH 663/829] chore: golf AlgebraicClosure construction using Isaacs (#19853) Also moves `Polynomial.isRoot_of_isRoot_iff_dvd_derivative_mul` to an earlier file. --- .../IsAlgClosed/AlgebraicClosure.lean | 326 ++---------------- Mathlib/FieldTheory/IsAlgClosed/Basic.lean | 27 ++ 2 files changed, 55 insertions(+), 298 deletions(-) diff --git a/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean b/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean index d91f6a62d0512..2c5ad01a7762a 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean @@ -3,10 +3,8 @@ Copyright (c) 2020 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ -import Mathlib.Algebra.DirectLimit import Mathlib.Algebra.CharP.Algebra -import Mathlib.Algebra.Polynomial.Eval.Irreducible -import Mathlib.FieldTheory.IsAlgClosed.Basic +import Mathlib.FieldTheory.Isaacs import Mathlib.FieldTheory.SplittingField.Construction /-! @@ -19,8 +17,9 @@ In this file we construct the algebraic closure of a field - `AlgebraicClosure k` is an algebraic closure of `k` (in the same universe). It is constructed by taking the polynomial ring generated by indeterminates `x_f` corresponding to monic irreducible polynomials `f` with coefficients in `k`, and quotienting - out by a maximal ideal containing every `f(x_f)`, and then repeating this step countably - many times. See Exercise 1.13 in Atiyah--Macdonald. + out by a maximal ideal containing every `f(x_f)`. Usually this needs to be repeated countably + many times (see Exercise 1.13 in [atiyah-macdonald]), but using a theorem of Isaacs [Isaacs1980], + once is enough (see https://kconrad.math.uconn.edu/blurbs/galoistheory/algclosure.pdf). ## Tags @@ -93,255 +92,14 @@ instance maxIdeal.isMaximal : (maxIdeal k).IsMaximal := theorem le_maxIdeal : spanEval k ≤ maxIdeal k := (Classical.choose_spec <| Ideal.exists_le_maximal _ <| spanEval_ne_top k).2 -/-- The first step of constructing `AlgebraicClosure`: adjoin a root of all monic polynomials -/ -def AdjoinMonic : Type u := - MvPolynomial (MonicIrreducible k) k ⧸ maxIdeal k - -instance AdjoinMonic.field : Field (AdjoinMonic k) := - Ideal.Quotient.field _ - -instance AdjoinMonic.inhabited : Inhabited (AdjoinMonic k) := - ⟨37⟩ - -/-- The canonical ring homomorphism to `AdjoinMonic k`. -/ -def toAdjoinMonic : k →+* AdjoinMonic k := - (Ideal.Quotient.mk _).comp C - -instance AdjoinMonic.algebra : Algebra k (AdjoinMonic k) := - (toAdjoinMonic k).toAlgebra - --- Porting note: In the statement, the type of `C` had to be made explicit. -theorem AdjoinMonic.algebraMap : algebraMap k (AdjoinMonic k) = (Ideal.Quotient.mk _).comp - (C : k →+* MvPolynomial (MonicIrreducible k) k) := rfl - -theorem AdjoinMonic.isIntegral (z : AdjoinMonic k) : IsIntegral k z := by - let ⟨p, hp⟩ := Ideal.Quotient.mk_surjective z - rw [← hp] - induction p using MvPolynomial.induction_on generalizing z with - | h_C => exact isIntegral_algebraMap - | h_add _ _ ha hb => exact (ha _ rfl).add (hb _ rfl) - | h_X p f ih => - refine @IsIntegral.mul k _ _ _ _ _ (Ideal.Quotient.mk (maxIdeal k) _) (ih _ rfl) ?_ - refine ⟨f, f.2.1, ?_⟩ - erw [AdjoinMonic.algebraMap, ← hom_eval₂, Ideal.Quotient.eq_zero_iff_mem] - exact le_maxIdeal k (Ideal.subset_span ⟨f, rfl⟩) - -theorem AdjoinMonic.exists_root {f : k[X]} (hfm : f.Monic) (hfi : Irreducible f) : - ∃ x : AdjoinMonic k, f.eval₂ (toAdjoinMonic k) x = 0 := - ⟨Ideal.Quotient.mk _ <| X (⟨f, hfm, hfi⟩ : MonicIrreducible k), by - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [toAdjoinMonic, ← hom_eval₂, Ideal.Quotient.eq_zero_iff_mem] - exact le_maxIdeal k (Ideal.subset_span <| ⟨_, rfl⟩)⟩ - -/-- The `n`th step of constructing `AlgebraicClosure`, together with its `Field` instance. -/ -def stepAux (n : ℕ) : Σ α : Type u, Field α := - Nat.recOn n ⟨k, inferInstance⟩ fun _ ih => ⟨@AdjoinMonic ih.1 ih.2, @AdjoinMonic.field ih.1 ih.2⟩ - -/-- The `n`th step of constructing `AlgebraicClosure`. -/ -def Step (n : ℕ) : Type u := - (stepAux k n).1 - --- Porting note: added during the port to help in the proof of `Step.isIntegral` below. -theorem Step.zero : Step k 0 = k := rfl - -instance Step.field (n : ℕ) : Field (Step k n) := - (stepAux k n).2 - --- Porting note: added during the port to help in the proof of `Step.isIntegral` below. -theorem Step.succ (n : ℕ) : Step k (n + 1) = AdjoinMonic (Step k n) := rfl - -instance Step.inhabited (n) : Inhabited (Step k n) := - ⟨37⟩ - -/-- The canonical inclusion to the `0`th step. -/ -def toStepZero : k →+* Step k 0 := - RingHom.id k - -/-- The canonical ring homomorphism to the next step. -/ -def toStepSucc (n : ℕ) : Step k n →+* (Step k (n + 1)) := - @toAdjoinMonic (Step k n) (Step.field k n) - -instance Step.algebraSucc (n) : Algebra (Step k n) (Step k (n + 1)) := - (toStepSucc k n).toAlgebra - -theorem toStepSucc.exists_root {n} {f : Polynomial (Step k n)} (hfm : f.Monic) - (hfi : Irreducible f) : ∃ x : Step k (n + 1), f.eval₂ (toStepSucc k n) x = 0 := - @AdjoinMonic.exists_root _ (Step.field k n) _ hfm hfi - --- Porting note: the following two declarations were added during the port to be used in the --- definition of toStepOfLE -private def toStepOfLE' (m n : ℕ) (h : m ≤ n) : Step k m → Step k n := -Nat.leRecOn h @fun a => toStepSucc k a - -private theorem toStepOfLE'.succ (m n : ℕ) (h : m ≤ n) : - toStepOfLE' k m (Nat.succ n) (h.trans n.le_succ) = - (toStepSucc k n) ∘ toStepOfLE' k m n h := by - ext x - exact Nat.leRecOn_succ h x - -/-- The canonical ring homomorphism to a step with a greater index. -/ -def toStepOfLE (m n : ℕ) (h : m ≤ n) : Step k m →+* Step k n where - toFun := toStepOfLE' k m n h - map_one' := by - induction' h with a h ih - · exact Nat.leRecOn_self 1 - · simp [toStepOfLE'.succ k m a h, ih] - map_mul' x y := by - simp only - induction' h with a h ih - · simp_rw [toStepOfLE', Nat.leRecOn_self] - · simp [toStepOfLE'.succ k m a h, ih] - map_zero' := by - simp only - induction' h with a h ih - · exact Nat.leRecOn_self 0 - · simp [toStepOfLE'.succ k m a h, ih] - map_add' x y := by - simp only - induction' h with a h ih - · simp_rw [toStepOfLE', Nat.leRecOn_self] - · simp [toStepOfLE'.succ k m a h, ih] - -@[simp] -theorem coe_toStepOfLE (m n : ℕ) (h : m ≤ n) : - (toStepOfLE k m n h : Step k m → Step k n) = Nat.leRecOn h @fun n => toStepSucc k n := - rfl - -instance Step.algebra (n) : Algebra k (Step k n) := - (toStepOfLE k 0 n n.zero_le).toAlgebra - -instance Step.scalar_tower (n) : IsScalarTower k (Step k n) (Step k (n + 1)) := - IsScalarTower.of_algebraMap_eq fun z => - @Nat.leRecOn_succ (Step k) 0 n n.zero_le (n + 1).zero_le (@fun n => toStepSucc k n) z - --- Porting note: Added to make `Step.isIntegral` faster -private theorem toStepOfLE.succ (n : ℕ) (h : 0 ≤ n) : - toStepOfLE k 0 (n + 1) (h.trans n.le_succ) = - (toStepSucc k n).comp (toStepOfLE k 0 n h) := by - ext1 x - rw [RingHom.comp_apply] - simp only [toStepOfLE, RingHom.coe_mk, MonoidHom.coe_mk, OneHom.coe_mk] - change _ = (_ ∘ _) x - rw [toStepOfLE'.succ k 0 n h] - -theorem Step.isIntegral (n) : ∀ z : Step k n, IsIntegral k z := by - induction' n with a h - · intro z - exact isIntegral_algebraMap - · intro z - change RingHom.IsIntegralElem _ _ - revert z - change RingHom.IsIntegral _ - unfold algebraMap - unfold Algebra.toRingHom - unfold algebra - unfold RingHom.toAlgebra - unfold RingHom.toAlgebra' - simp only - rw [toStepOfLE.succ k a a.zero_le] - apply @RingHom.IsIntegral.trans (Step k 0) (Step k a) (Step k (a + 1)) _ _ _ - (toStepOfLE k 0 a (a.zero_le : 0 ≤ a)) (toStepSucc k a) _ - · intro z - convert AdjoinMonic.isIntegral (Step k a) (z : Step k (a + 1)) - · convert h -- Porting note: This times out at 500000 - -instance toStepOfLE.directedSystem : DirectedSystem (Step k) fun i j h => toStepOfLE k i j h := - ⟨fun _ => Nat.leRecOn_self, fun _ _ _ h₁₂ h₂₃ x => (Nat.leRecOn_trans h₁₂ h₂₃ x).symm⟩ - end AlgebraicClosure -/-- Auxiliary construction for `AlgebraicClosure`. Although `AlgebraicClosureAux` does define -the algebraic closure of a field, it is redefined at `AlgebraicClosure` in order to make sure -certain instance diamonds commute by definition. --/ -def AlgebraicClosureAux [Field k] : Type u := - Ring.DirectLimit (AlgebraicClosure.Step k) fun i j h => AlgebraicClosure.toStepOfLE k i j h - -namespace AlgebraicClosureAux - -open AlgebraicClosure - -/-- `AlgebraicClosureAux k` is a `Field` -/ -local instance field : Field (AlgebraicClosureAux k) := - Field.DirectLimit.field _ _ - -instance : Inhabited (AlgebraicClosureAux k) := - ⟨37⟩ - -/-- The canonical ring embedding from the `n`th step to the algebraic closure. -/ -def ofStep (n : ℕ) : Step k n →+* AlgebraicClosureAux k := - Ring.DirectLimit.of _ _ _ - -theorem ofStep_succ (n : ℕ) : (ofStep k (n + 1)).comp (toStepSucc k n) = ofStep k n := by - ext x - have hx : toStepOfLE' k n (n+1) n.le_succ x = toStepSucc k n x := Nat.leRecOn_succ' x - unfold ofStep - rw [RingHom.comp_apply] - dsimp [toStepOfLE] - rw [← hx] - change Ring.DirectLimit.of (Step k) (toStepOfLE' k) (n + 1) (_) = - Ring.DirectLimit.of (Step k) (toStepOfLE' k) n x - convert Ring.DirectLimit.of_f n.le_succ x - -- Porting Note: Original proof timed out at 2 mil. Heartbeats. The problem was likely - -- in comparing `toStepOfLE'` with `toStepSucc`. In the above, I made some things more explicit - -- Original proof: - -- RingHom.ext fun x => - -- show Ring.DirectLimit.of (Step k) (fun i j h => toStepOfLE k i j h) _ _ = _ by - -- convert Ring.DirectLimit.of_f n.le_succ x; ext x; exact (Nat.leRecOn_succ' x).symm - -theorem exists_ofStep (z : AlgebraicClosureAux k) : ∃ n x, ofStep k n x = z := - Ring.DirectLimit.exists_of z - -theorem exists_root {f : Polynomial (AlgebraicClosureAux k)} - (hfm : f.Monic) (hfi : Irreducible f) : ∃ x : AlgebraicClosureAux k, f.eval x = 0 := by - have : ∃ n p, Polynomial.map (ofStep k n) p = f := by - convert Ring.DirectLimit.Polynomial.exists_of f - obtain ⟨n, p, rfl⟩ := this - rw [monic_map_iff] at hfm - have := hfm.irreducible_of_irreducible_map (ofStep k n) p hfi - obtain ⟨x, hx⟩ := toStepSucc.exists_root k hfm this - refine ⟨ofStep k (n + 1) x, ?_⟩ - rw [← ofStep_succ k n, eval_map, ← hom_eval₂, hx, RingHom.map_zero] - -@[local instance] theorem instIsAlgClosed : IsAlgClosed (AlgebraicClosureAux k) := - IsAlgClosed.of_exists_root _ fun _ => exists_root k - -/-- `AlgebraicClosureAux k` is a `k`-`Algebra` -/ -local instance instAlgebra : Algebra k (AlgebraicClosureAux k) := - (ofStep k 0).toAlgebra - -/-- Canonical algebra embedding from the `n`th step to the algebraic closure. -/ -def ofStepHom (n) : Step k n →ₐ[k] AlgebraicClosureAux k := - { ofStep k n with - commutes' := by - -- Porting note: Originally `(fun x => Ring.DirectLimit.of_f n.zero_le x)` - -- I think one problem was in recognizing that we want `toStepOfLE` in `of_f` - intro x - simp only [RingHom.toMonoidHom_eq_coe, OneHom.toFun_eq_coe, MonoidHom.toOneHom_coe, - MonoidHom.coe_coe] - convert @Ring.DirectLimit.of_f ℕ _ (Step k) _ (fun m n h => (toStepOfLE k m n h : _ → _)) - 0 n n.zero_le x } - -instance isAlgebraic : Algebra.IsAlgebraic k (AlgebraicClosureAux k) := - ⟨fun z => - IsIntegral.isAlgebraic <| - let ⟨n, x, hx⟩ := exists_ofStep k z - hx ▸ (Step.isIntegral k n x).map (ofStepHom k n)⟩ - -@[local instance] theorem isAlgClosure : IsAlgClosure k (AlgebraicClosureAux k) := - ⟨AlgebraicClosureAux.instIsAlgClosed k, isAlgebraic k⟩ - -end AlgebraicClosureAux - -attribute [local instance] AlgebraicClosureAux.field AlgebraicClosureAux.instAlgebra - AlgebraicClosureAux.instIsAlgClosed - +open AlgebraicClosure in /-- The canonical algebraic closure of a field, the direct limit of adding roots to the field for each polynomial over the field. -/ @[stacks 09GT] def AlgebraicClosure : Type u := - MvPolynomial (AlgebraicClosureAux k) k ⧸ - RingHom.ker (MvPolynomial.aeval (R := k) id).toRingHom + MvPolynomial (MonicIrreducible k) k ⧸ maxIdeal k namespace AlgebraicClosure @@ -358,21 +116,8 @@ instance {R S : Type*} [CommSemiring R] [CommSemiring S] [Algebra R S] [Algebra [IsScalarTower R S k] : IsScalarTower R S (AlgebraicClosure k) := Ideal.Quotient.isScalarTower _ _ _ -/-- The equivalence between `AlgebraicClosure` and `AlgebraicClosureAux`, which we use to transfer -properties of `AlgebraicClosureAux` to `AlgebraicClosure` -/ -def algEquivAlgebraicClosureAux : - AlgebraicClosure k ≃ₐ[k] AlgebraicClosureAux k := by - delta AlgebraicClosure - exact Ideal.quotientKerAlgEquivOfSurjective - (fun x => ⟨MvPolynomial.X x, by simp⟩) - --- Those two instances are copy-pasta from the analogous instances for `SplittingField` -instance instGroupWithZero : GroupWithZero (AlgebraicClosure k) := - let e := algEquivAlgebraicClosureAux k - { inv := fun a ↦ e.symm (e a)⁻¹ - inv_zero := by simp - mul_inv_cancel := fun a ha ↦ e.injective <| by simp [EmbeddingLike.map_ne_zero_iff.2 ha] - __ := e.surjective.nontrivial } +instance instGroupWithZero : GroupWithZero (AlgebraicClosure k) where + __ := Ideal.Quotient.field _ instance instField : Field (AlgebraicClosure k) where __ := instCommRing _ @@ -389,15 +134,27 @@ instance instField : Field (AlgebraicClosure k) where qsmul_def q x := Quotient.inductionOn x fun p ↦ congr_arg Quotient.mk'' <| by ext; simp [MvPolynomial.algebraMap_eq, Rat.smul_def] -instance isAlgClosed : IsAlgClosed (AlgebraicClosure k) := - IsAlgClosed.of_ringEquiv _ _ (algEquivAlgebraicClosureAux k).symm.toRingEquiv - -instance : IsAlgClosure k (AlgebraicClosure k) := by - rw [isAlgClosure_iff] - exact ⟨inferInstance, (algEquivAlgebraicClosureAux k).symm.isAlgebraic⟩ - instance isAlgebraic : Algebra.IsAlgebraic k (AlgebraicClosure k) := - IsAlgClosure.isAlgebraic + ⟨fun z => + IsIntegral.isAlgebraic <| by + let ⟨p, hp⟩ := Ideal.Quotient.mk_surjective z + rw [← hp] + induction p using MvPolynomial.induction_on generalizing z with + | h_C => exact isIntegral_algebraMap + | h_add _ _ ha hb => exact (ha _ rfl).add (hb _ rfl) + | h_X p f ih => + refine @IsIntegral.mul k _ _ _ _ _ (Ideal.Quotient.mk (maxIdeal k) _) (ih _ rfl) ?_ + refine ⟨f, f.2.1, ?_⟩ + erw [algebraMap, ← hom_eval₂, Ideal.Quotient.eq_zero_iff_mem] + exact le_maxIdeal k (Ideal.subset_span ⟨f, rfl⟩)⟩ + +instance : IsAlgClosure k (AlgebraicClosure k) := .of_exist_roots fun f hfm hfi ↦ + ⟨Ideal.Quotient.mk _ <| MvPolynomial.X (⟨f, hfm, hfi⟩ : MonicIrreducible k), by + -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 + erw [aeval_def, ← hom_eval₂, Ideal.Quotient.eq_zero_iff_mem] + exact le_maxIdeal k (Ideal.subset_span <| ⟨_, rfl⟩)⟩ + +instance isAlgClosed : IsAlgClosed (AlgebraicClosure k) := IsAlgClosure.isAlgClosed k instance [CharZero k] : CharZero (AlgebraicClosure k) := charZero_of_injective_algebraMap (RingHom.injective (algebraMap k (AlgebraicClosure k))) @@ -411,30 +168,3 @@ instance {L : Type*} [Field k] [Field L] [Algebra k L] [Algebra.IsAlgebraic k L] isAlgClosed := inferInstance end AlgebraicClosure - -/-- Over an algebraically closed field of characteristic zero a necessary and sufficient condition -for the set of roots of a nonzero polynomial `f` to be a subset of the set of roots of `g` is that -`f` divides `f.derivative * g`. Over an integral domain, this is a sufficient but not necessary -condition. See `isRoot_of_isRoot_of_dvd_derivative_mul` -/ -theorem Polynomial.isRoot_of_isRoot_iff_dvd_derivative_mul {K : Type*} [Field K] - [IsAlgClosed K] [CharZero K] {f g : K[X]} (hf0 : f ≠ 0) : - (∀ x, IsRoot f x → IsRoot g x) ↔ f ∣ f.derivative * g := by - refine ⟨?_, isRoot_of_isRoot_of_dvd_derivative_mul hf0⟩ - by_cases hg0 : g = 0 - · simp [hg0] - by_cases hdf0 : derivative f = 0 - · rw [eq_C_of_derivative_eq_zero hdf0] - simp only [eval_C, derivative_C, zero_mul, dvd_zero, implies_true] - have hdg : f.derivative * g ≠ 0 := mul_ne_zero hdf0 hg0 - classical rw [Splits.dvd_iff_roots_le_roots (IsAlgClosed.splits f) hf0 hdg, Multiset.le_iff_count] - simp only [count_roots, rootMultiplicity_mul hdg] - refine forall_imp fun a => ?_ - by_cases haf : f.eval a = 0 - · have h0 : 0 < f.rootMultiplicity a := (rootMultiplicity_pos hf0).2 haf - rw [derivative_rootMultiplicity_of_root haf] - intro h - calc rootMultiplicity a f - = rootMultiplicity a f - 1 + 1 := (Nat.sub_add_cancel (Nat.succ_le_iff.1 h0)).symm - _ ≤ rootMultiplicity a f - 1 + rootMultiplicity a g := add_le_add le_rfl (Nat.succ_le_iff.1 - ((rootMultiplicity_pos hg0).2 (h haf))) - · simp [haf, rootMultiplicity_eq_zero haf] diff --git a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean index c164bdbac9521..02cc7fb943f06 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean @@ -497,3 +497,30 @@ theorem Algebra.IsAlgebraic.algHomEquivAlgHomOfSplits_apply_apply (L : Type*) [F Algebra.IsAlgebraic.algHomEquivAlgHomOfSplits A L hL f x = algebraMap L A (f x) := rfl end Algebra.IsAlgebraic + +/-- Over an algebraically closed field of characteristic zero a necessary and sufficient condition +for the set of roots of a nonzero polynomial `f` to be a subset of the set of roots of `g` is that +`f` divides `f.derivative * g`. Over an integral domain, this is a sufficient but not necessary +condition. See `isRoot_of_isRoot_of_dvd_derivative_mul` -/ +theorem Polynomial.isRoot_of_isRoot_iff_dvd_derivative_mul {K : Type*} [Field K] + [IsAlgClosed K] [CharZero K] {f g : K[X]} (hf0 : f ≠ 0) : + (∀ x, IsRoot f x → IsRoot g x) ↔ f ∣ f.derivative * g := by + refine ⟨?_, isRoot_of_isRoot_of_dvd_derivative_mul hf0⟩ + by_cases hg0 : g = 0 + · simp [hg0] + by_cases hdf0 : derivative f = 0 + · rw [eq_C_of_derivative_eq_zero hdf0] + simp only [eval_C, derivative_C, zero_mul, dvd_zero, implies_true] + have hdg : f.derivative * g ≠ 0 := mul_ne_zero hdf0 hg0 + classical rw [Splits.dvd_iff_roots_le_roots (IsAlgClosed.splits f) hf0 hdg, Multiset.le_iff_count] + simp only [count_roots, rootMultiplicity_mul hdg] + refine forall_imp fun a => ?_ + by_cases haf : f.eval a = 0 + · have h0 : 0 < f.rootMultiplicity a := (rootMultiplicity_pos hf0).2 haf + rw [derivative_rootMultiplicity_of_root haf] + intro h + calc rootMultiplicity a f + = rootMultiplicity a f - 1 + 1 := (Nat.sub_add_cancel (Nat.succ_le_iff.1 h0)).symm + _ ≤ rootMultiplicity a f - 1 + rootMultiplicity a g := add_le_add le_rfl (Nat.succ_le_iff.1 + ((rootMultiplicity_pos hg0).2 (h haf))) + · simp [haf, rootMultiplicity_eq_zero haf] From 2c66a57507c70c129eb05a900301d73b3e340918 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 11 Dec 2024 14:33:39 +0000 Subject: [PATCH 664/829] feat(Order/SuccPred/Limit): extra lemmas (#19801) These are trivial variations on already proven results. --- Mathlib/Order/SuccPred/Limit.lean | 28 ++++++++++++++++++++++++++++ 1 file changed, 28 insertions(+) diff --git a/Mathlib/Order/SuccPred/Limit.lean b/Mathlib/Order/SuccPred/Limit.lean index 6170a1e80ff57..2497ef7e19efd 100644 --- a/Mathlib/Order/SuccPred/Limit.lean +++ b/Mathlib/Order/SuccPred/Limit.lean @@ -314,10 +314,24 @@ theorem IsSuccPrelimit.le_iff_forall_le (h : IsSuccPrelimit a) : a ≤ b ↔ ∀ by_contra! ha exact h b ⟨ha, fun c hb hc ↦ (H c hc).not_lt hb⟩ +theorem IsSuccLimit.le_iff_forall_le (h : IsSuccLimit a) : a ≤ b ↔ ∀ c < a, c ≤ b := + h.isSuccPrelimit.le_iff_forall_le + theorem IsSuccPrelimit.lt_iff_exists_lt (h : IsSuccPrelimit b) : a < b ↔ ∃ c < b, a < c := by rw [← not_iff_not] simp [h.le_iff_forall_le] +theorem IsSuccLimit.lt_iff_exists_lt (h : IsSuccLimit b) : a < b ↔ ∃ c < b, a < c := + h.isSuccPrelimit.lt_iff_exists_lt + +variable [SuccOrder α] + +theorem IsSuccPrelimit.le_succ_iff (hb : IsSuccPrelimit b) : b ≤ succ a ↔ b ≤ a := + le_iff_le_iff_lt_iff_lt.2 hb.succ_lt_iff + +theorem IsSuccLimit.le_succ_iff (hb : IsSuccLimit b) : b ≤ succ a ↔ b ≤ a := + hb.isSuccPrelimit.le_succ_iff + end LinearOrder /-! ### Predecessor limits -/ @@ -597,9 +611,23 @@ variable [LinearOrder α] theorem IsPredPrelimit.le_iff_forall_le (h : IsPredPrelimit a) : b ≤ a ↔ ∀ ⦃c⦄, a < c → b ≤ c := h.dual.le_iff_forall_le +theorem IsPredLimit.le_iff_forall_le (h : IsPredLimit a) : b ≤ a ↔ ∀ ⦃c⦄, a < c → b ≤ c := + h.dual.le_iff_forall_le + theorem IsPredPrelimit.lt_iff_exists_lt (h : IsPredPrelimit b) : b < a ↔ ∃ c, b < c ∧ c < a := h.dual.lt_iff_exists_lt +theorem IsPredLimit.lt_iff_exists_lt (h : IsPredLimit b) : b < a ↔ ∃ c, b < c ∧ c < a := + h.dual.lt_iff_exists_lt + +variable [PredOrder α] + +theorem IsPredPrelimit.pred_le_iff (hb : IsPredPrelimit b) : pred a ≤ b ↔ a ≤ b := + hb.dual.le_succ_iff + +theorem IsPredLimit.pred_le_iff (hb : IsPredLimit b) : pred a ≤ b ↔ a ≤ b := + hb.dual.le_succ_iff + end LinearOrder end Order From e52bfb49fec3da26a2acc8798a051ceb88cd0692 Mon Sep 17 00:00:00 2001 From: Scott Carnahan <128885296+ScottCarnahan@users.noreply.github.com> Date: Wed, 11 Dec 2024 16:24:23 +0000 Subject: [PATCH 665/829] chore(RingTheory/Binomial): make a variable implicit, misc cleanups (#19856) This PR changes the `Nat` variable in `nsmul_right_injective` to implicit. Also, the `BinomialRing` instance on `Nat` is changed to use `Nat.multichoose`, since (contrary to my previous impression) that function seems unlikely to be deprecated. --- Mathlib/NumberTheory/Padics/MahlerBasis.lean | 2 +- Mathlib/RingTheory/Binomial.lean | 55 ++++++++++---------- 2 files changed, 28 insertions(+), 29 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/MahlerBasis.lean b/Mathlib/NumberTheory/Padics/MahlerBasis.lean index 827e9456b8035..55ebe61b44164 100644 --- a/Mathlib/NumberTheory/Padics/MahlerBasis.lean +++ b/Mathlib/NumberTheory/Padics/MahlerBasis.lean @@ -68,7 +68,7 @@ lemma norm_ascPochhammer_le (k : ℕ) (x : ℤ_[p]) : /-- The p-adic integers are a binomial ring, i.e. a ring where binomial coefficients make sense. -/ noncomputable instance instBinomialRing : BinomialRing ℤ_[p] where - nsmul_right_injective n hn := smul_right_injective ℤ_[p] hn + nsmul_right_injective hn := smul_right_injective ℤ_[p] hn -- We define `multichoose` as a fraction in `ℚ_[p]` together with a proof that its norm is `≤ 1`. multichoose x k := ⟨(ascPochhammer ℤ_[p] k).eval x / (k.factorial : ℚ_[p]), by rw [norm_div, div_le_one (by simpa using k.factorial_ne_zero)] diff --git a/Mathlib/RingTheory/Binomial.lean b/Mathlib/RingTheory/Binomial.lean index ad5a0893a9a04..dc1ee22987949 100644 --- a/Mathlib/RingTheory/Binomial.lean +++ b/Mathlib/RingTheory/Binomial.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Carnahan -/ import Mathlib.Algebra.Polynomial.Smeval -import Mathlib.Algebra.Order.Floor import Mathlib.GroupTheory.GroupAction.Ring import Mathlib.RingTheory.Polynomial.Pochhammer import Mathlib.Tactic.FieldSimp @@ -67,7 +66,7 @@ number powers, but retain the ring name. We introduce `Ring.multichoose` as the quotient. -/ class BinomialRing (R : Type*) [AddCommMonoid R] [Pow R ℕ] where /-- Scalar multiplication by positive integers is injective -/ - nsmul_right_injective (n : ℕ) (h : n ≠ 0) : Injective (n • · : R → R) + nsmul_right_injective {n : ℕ} (h : n ≠ 0) : Injective (n • · : R → R) /-- A multichoose function, giving the quotient of Pochhammer evaluations by factorials. -/ multichoose : R → ℕ → R /-- The `n`th ascending Pochhammer polynomial evaluated at any element is divisible by `n!` -/ @@ -78,8 +77,8 @@ namespace Ring variable {R : Type*} [AddCommMonoid R] [Pow R ℕ] [BinomialRing R] -theorem nsmul_right_injective (n : ℕ) (h : n ≠ 0) : - Injective (n • · : R → R) := BinomialRing.nsmul_right_injective n h +theorem nsmul_right_injective {n : ℕ} (h : n ≠ 0) : + Injective (n • · : R → R) := BinomialRing.nsmul_right_injective h /-- The multichoose function is the quotient of ascending Pochhammer evaluation by the corresponding factorial. When applied to natural numbers, `multichoose k n` describes choosing a multiset of `n` @@ -96,7 +95,7 @@ theorem factorial_nsmul_multichoose_eq_ascPochhammer (r : R) (n : ℕ) : @[simp] theorem multichoose_zero_right' (r : R) : multichoose r 0 = r ^ 0 := by - refine nsmul_right_injective (Nat.factorial 0) (Nat.factorial_ne_zero 0) ?_ + refine nsmul_right_injective (Nat.factorial_ne_zero 0) ?_ simp only rw [factorial_nsmul_multichoose_eq_ascPochhammer, ascPochhammer_zero, smeval_one, Nat.factorial] @@ -106,7 +105,7 @@ theorem multichoose_zero_right [MulOneClass R] [NatPowAssoc R] @[simp] theorem multichoose_one_right' (r : R) : multichoose r 1 = r ^ 1 := by - refine nsmul_right_injective (Nat.factorial 1) (Nat.factorial_ne_zero 1) ?_ + refine nsmul_right_injective (Nat.factorial_ne_zero 1) ?_ simp only rw [factorial_nsmul_multichoose_eq_ascPochhammer, ascPochhammer_one, smeval_X, Nat.factorial_one, one_smul] @@ -118,7 +117,7 @@ variable {R : Type*} [NonAssocSemiring R] [Pow R ℕ] [NatPowAssoc R] [BinomialR @[simp] theorem multichoose_zero_succ (k : ℕ) : multichoose (0 : R) (k + 1) = 0 := by - refine nsmul_right_injective (Nat.factorial (k + 1)) (Nat.factorial_ne_zero (k + 1)) ?_ + refine nsmul_right_injective (Nat.factorial_ne_zero (k + 1)) ?_ simp only rw [factorial_nsmul_multichoose_eq_ascPochhammer, smul_zero, ascPochhammer_succ_left, smeval_X_mul, zero_mul] @@ -127,7 +126,7 @@ theorem ascPochhammer_succ_succ (r : R) (k : ℕ) : smeval (ascPochhammer ℕ (k + 1)) (r + 1) = Nat.factorial (k + 1) • multichoose (r + 1) k + smeval (ascPochhammer ℕ (k + 1)) r := by nth_rw 1 [ascPochhammer_succ_right, ascPochhammer_succ_left, mul_comm (ascPochhammer ℕ k)] - simp only [smeval_mul, smeval_comp ℕ _ _ r, smeval_add, smeval_X] + simp only [smeval_mul, smeval_comp, smeval_add, smeval_X] rw [Nat.factorial, mul_smul, factorial_nsmul_multichoose_eq_ascPochhammer] simp only [smeval_one, npow_one, npow_zero, one_smul] rw [← C_eq_natCast, smeval_C, npow_zero, add_assoc, add_mul, add_comm 1, @nsmul_one, add_mul] @@ -136,7 +135,7 @@ theorem ascPochhammer_succ_succ (r : R) (k : ℕ) : theorem multichoose_succ_succ (r : R) (k : ℕ) : multichoose (r + 1) (k + 1) = multichoose r (k + 1) + multichoose (r + 1) k := by - refine nsmul_right_injective (Nat.factorial (k + 1)) (Nat.factorial_ne_zero (k + 1)) ?_ + refine nsmul_right_injective (Nat.factorial_ne_zero (k + 1)) ?_ simp only [factorial_nsmul_multichoose_eq_ascPochhammer, smul_add] rw [add_comm (smeval (ascPochhammer ℕ (k+1)) r), ascPochhammer_succ_succ r k] @@ -225,10 +224,10 @@ section Basic_Instances open Polynomial instance Nat.instBinomialRing : BinomialRing ℕ where - nsmul_right_injective _ hn _ _ hrs := Nat.eq_of_mul_eq_mul_left (Nat.pos_of_ne_zero hn) hrs - multichoose n k := Nat.choose (n + k - 1) k + nsmul_right_injective hn _ _ hrs := Nat.eq_of_mul_eq_mul_left (Nat.pos_of_ne_zero hn) hrs + multichoose := Nat.multichoose factorial_nsmul_multichoose r n := by - rw [smul_eq_mul, ← Nat.descFactorial_eq_factorial_mul_choose, + rw [smul_eq_mul, Nat.multichoose_eq r n, ← Nat.descFactorial_eq_factorial_mul_choose, ← eval_eq_smeval r (ascPochhammer ℕ n), ascPochhammer_nat_eq_descFactorial] /-- The multichoose function for integers. -/ @@ -238,7 +237,7 @@ def Int.multichoose (n : ℤ) (k : ℕ) : ℤ := | negSucc n => (-1) ^ k * Nat.choose (n + 1) k instance Int.instBinomialRing : BinomialRing ℤ where - nsmul_right_injective _ hn _ _ hrs := Int.eq_of_mul_eq_mul_left (Int.ofNat_ne_zero.mpr hn) hrs + nsmul_right_injective hn _ _ hrs := Int.eq_of_mul_eq_mul_left (Int.ofNat_ne_zero.mpr hn) hrs multichoose := Int.multichoose factorial_nsmul_multichoose r k := by rw [Int.multichoose.eq_def, nsmul_eq_mul] @@ -254,7 +253,7 @@ instance Int.instBinomialRing : BinomialRing ℤ where ← Int.neg_ofNat_succ, ascPochhammer_smeval_neg_eq_descPochhammer] noncomputable instance {R : Type*} [AddCommMonoid R] [Module ℚ≥0 R] [Pow R ℕ] : BinomialRing R where - nsmul_right_injective n hn r s hrs := by + nsmul_right_injective {n} hn r s hrs := by rw [← one_smul ℚ≥0 r, ← one_smul ℚ≥0 s, show 1 = (n : ℚ≥0)⁻¹ • (n : ℚ≥0) by simp_all] simp_all only [smul_assoc, Nat.cast_smul_eq_nsmul] multichoose r n := (n.factorial : ℚ≥0)⁻¹ • Polynomial.smeval (ascPochhammer ℕ n) r @@ -309,7 +308,7 @@ theorem smeval_ascPochhammer_nat_cast {R} [NonAssocRing R] [Pow R ℕ] [NatPowAs rw [smeval_at_natCast (ascPochhammer ℕ k) n] theorem multichoose_neg_self (n : ℕ) : multichoose (-n : ℤ) n = (-1)^n := by - apply nsmul_right_injective _ (Nat.factorial_ne_zero _) + apply nsmul_right_injective (Nat.factorial_ne_zero _) on_goal 1 => simp only -- This closes both remaining goals at once. rw [factorial_nsmul_multichoose_eq_ascPochhammer, smeval_ascPochhammer_self_neg, nsmul_eq_mul, @@ -317,25 +316,25 @@ theorem multichoose_neg_self (n : ℕ) : multichoose (-n : ℤ) n = (-1)^n := by @[simp] theorem multichoose_neg_succ (n : ℕ) : multichoose (-n : ℤ) (n + 1) = 0 := by - apply nsmul_right_injective _ (Nat.factorial_ne_zero _) + apply nsmul_right_injective (Nat.factorial_ne_zero _) on_goal 1 => simp only -- This closes both remaining goals at once. rw [factorial_nsmul_multichoose_eq_ascPochhammer, smeval_ascPochhammer_succ_neg, smul_zero] theorem multichoose_neg_add (n k : ℕ) : multichoose (-n : ℤ) (n + k + 1) = 0 := by - refine nsmul_right_injective (Nat.factorial (n + k + 1)) (Nat.factorial_ne_zero (n + k + 1)) ?_ + refine nsmul_right_injective (Nat.factorial_ne_zero (n + k + 1)) ?_ simp only rw [factorial_nsmul_multichoose_eq_ascPochhammer, smeval_ascPochhammer_neg_add, smul_zero] @[simp] theorem multichoose_neg_of_lt (n k : ℕ) (h : n < k) : multichoose (-n : ℤ) k = 0 := by - refine nsmul_right_injective (Nat.factorial k) (Nat.factorial_ne_zero k) ?_ + refine nsmul_right_injective (Nat.factorial_ne_zero k) ?_ simp only rw [factorial_nsmul_multichoose_eq_ascPochhammer, smeval_ascPochhammer_neg_of_lt h, smul_zero] theorem multichoose_succ_neg_natCast [NatPowAssoc R] (n : ℕ) : multichoose (-n : R) (n + 1) = 0 := by - refine nsmul_right_injective (Nat.factorial (n + 1)) (Nat.factorial_ne_zero (n + 1)) ?_ + refine nsmul_right_injective (Nat.factorial_ne_zero (n + 1)) ?_ simp only [smul_zero] rw [factorial_nsmul_multichoose_eq_ascPochhammer, smeval_neg_nat, smeval_ascPochhammer_succ_neg n, Int.cast_zero] @@ -381,7 +380,7 @@ theorem descPochhammer_eq_factorial_smul_choose [NatPowAssoc R] (r : R) (n : ℕ rw [h, ascPochhammer_smeval_cast, add_comm_sub] theorem choose_natCast [NatPowAssoc R] (n k : ℕ) : choose (n : R) k = Nat.choose n k := by - refine nsmul_right_injective (Nat.factorial k) (Nat.factorial_ne_zero k) ?_ + refine nsmul_right_injective (Nat.factorial_ne_zero k) ?_ simp only rw [← descPochhammer_eq_factorial_smul_choose, nsmul_eq_mul, ← Nat.cast_mul, ← Nat.descFactorial_eq_factorial_mul_choose, ← descPochhammer_smeval_eq_descFactorial] @@ -392,7 +391,7 @@ alias choose_nat_cast := choose_natCast @[simp] theorem choose_zero_right' (r : R) : choose r 0 = (r + 1) ^ 0 := by dsimp only [choose] - refine nsmul_right_injective (Nat.factorial 0) (Nat.factorial_ne_zero 0) ?_ + refine nsmul_right_injective (Nat.factorial_ne_zero 0) ?_ simp [factorial_nsmul_multichoose_eq_ascPochhammer] theorem choose_zero_right [NatPowAssoc R] (r : R) : choose r 0 = 1 := by @@ -425,15 +424,15 @@ theorem descPochhammer_succ_succ_smeval {R} [NonAssocRing R] [Pow R ℕ] [NatPow (k + 1) • smeval (descPochhammer ℤ k) r + smeval (descPochhammer ℤ (k + 1)) r := by nth_rw 1 [descPochhammer_succ_left] rw [descPochhammer_succ_right, mul_comm (descPochhammer ℤ k)] - simp only [smeval_comp ℤ _ _ (r + 1), smeval_sub, smeval_add, smeval_mul, smeval_X, smeval_one, - npow_one, npow_zero, one_smul, add_sub_cancel_right, sub_mul, add_mul, add_smul, one_mul] + simp only [smeval_comp, smeval_sub, smeval_add, smeval_mul, smeval_X, smeval_one, npow_one, + npow_zero, one_smul, add_sub_cancel_right, sub_mul, add_mul, add_smul, one_mul] rw [← C_eq_natCast, smeval_C, npow_zero, add_comm (k • smeval (descPochhammer ℤ k) r) _, add_assoc, add_comm (k • smeval (descPochhammer ℤ k) r) _, ← add_assoc, ← add_sub_assoc, nsmul_eq_mul, zsmul_one, Int.cast_natCast, sub_add_cancel, add_comm] theorem choose_succ_succ [NatPowAssoc R] (r : R) (k : ℕ) : choose (r + 1) (k + 1) = choose r k + choose r (k + 1) := by - refine nsmul_right_injective (Nat.factorial (k + 1)) (Nat.factorial_ne_zero (k + 1)) ?_ + refine nsmul_right_injective (Nat.factorial_ne_zero (k + 1)) ?_ simp only [smul_add, ← descPochhammer_eq_factorial_smul_choose] rw [Nat.factorial_succ, mul_smul, ← descPochhammer_eq_factorial_smul_choose r, descPochhammer_succ_succ_smeval r k] @@ -447,9 +446,9 @@ theorem choose_eq_nat_choose [NatPowAssoc R] (n k : ℕ) : choose (n : R) k = Na | zero => rw [choose_zero_right, Nat.choose_zero_right, Nat.cast_one] | succ k => rw [Nat.cast_succ, choose_succ_succ, ih, ih, Nat.choose_succ_succ, Nat.cast_add] -theorem choose_smul_choose [NatPowAssoc R] (r : R) (n k : ℕ) (hkn : k ≤ n) : +theorem choose_smul_choose [NatPowAssoc R] (r : R) {n k : ℕ} (hkn : k ≤ n) : (Nat.choose n k) • choose r n = choose r k * choose (r - k) (n - k) := by - refine nsmul_right_injective (Nat.factorial n) (Nat.factorial_ne_zero n) ?_ + refine nsmul_right_injective (Nat.factorial_ne_zero n) ?_ simp only rw [nsmul_left_comm, ← descPochhammer_eq_factorial_smul_choose, ← Nat.choose_mul_factorial_mul_factorial hkn, ← smul_mul_smul_comm, @@ -461,7 +460,7 @@ theorem choose_smul_choose [NatPowAssoc R] (r : R) (n k : ℕ) (hkn : k ≤ n) : theorem choose_add_smul_choose [NatPowAssoc R] (r : R) (n k : ℕ) : (Nat.choose (n + k) k) • choose (r + k) (n + k) = choose (r + k) k * choose r n := by - rw [choose_smul_choose (r + k) (n + k) k (Nat.le_add_left k n), Nat.add_sub_cancel, + rw [choose_smul_choose (r + k) (Nat.le_add_left k n), Nat.add_sub_cancel, add_sub_cancel_right] end @@ -501,7 +500,7 @@ theorem descPochhammer_smeval_add [Ring R] {r s : R} (k : ℕ) (h: Commute r s) theorem add_choose_eq [Ring R] [BinomialRing R] {r s : R} (k : ℕ) (h : Commute r s) : choose (r + s) k = ∑ ij ∈ antidiagonal k, choose r ij.1 * choose s ij.2 := by - refine nsmul_right_injective (Nat.factorial k) (Nat.factorial_ne_zero k) ?_ + refine nsmul_right_injective (Nat.factorial_ne_zero k) ?_ simp only rw [← descPochhammer_eq_factorial_smul_choose, smul_sum, descPochhammer_smeval_add _ h] refine sum_congr rfl ?_ From fbe88317752b5107a35d7d031c0590efe6a8a0bb Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Wed, 11 Dec 2024 17:32:05 +0000 Subject: [PATCH 666/829] feat(Probability): parallel composition of kernels (#19482) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Parallel composition of kernels: from `κ : Kernel α β` and `η : Kernel γ δ`, define a kernel `κ ∥ₖ η` from `α × γ` to `β × δ`: `(κ ∥ₖ η) (a, c) = (κ a).prod (η c)`. From the TestingLowerBounds project. Co-authored-by: Lorenzo Luccioli Co-authored-by: Rémy Degenne --- Mathlib.lean | 1 + .../Probability/Kernel/Composition/Basic.lean | 12 ++ .../Kernel/Composition/ParallelComp.lean | 118 ++++++++++++++++++ 3 files changed, 131 insertions(+) create mode 100644 Mathlib/Probability/Kernel/Composition/ParallelComp.lean diff --git a/Mathlib.lean b/Mathlib.lean index 8569d32bdc872..673308d1c2f37 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4179,6 +4179,7 @@ import Mathlib.Probability.Kernel.Basic import Mathlib.Probability.Kernel.Composition.Basic import Mathlib.Probability.Kernel.Composition.IntegralCompProd import Mathlib.Probability.Kernel.Composition.MeasureCompProd +import Mathlib.Probability.Kernel.Composition.ParallelComp import Mathlib.Probability.Kernel.CondDistrib import Mathlib.Probability.Kernel.Condexp import Mathlib.Probability.Kernel.Defs diff --git a/Mathlib/Probability/Kernel/Composition/Basic.lean b/Mathlib/Probability/Kernel/Composition/Basic.lean index e3b9daa87e50a..a992511575a5d 100644 --- a/Mathlib/Probability/Kernel/Composition/Basic.lean +++ b/Mathlib/Probability/Kernel/Composition/Basic.lean @@ -726,6 +726,12 @@ instance IsMarkovKernel.comap (κ : Kernel α β) [IsMarkovKernel κ] (hg : Meas IsMarkovKernel (comap κ g hg) := ⟨fun a => ⟨by rw [comap_apply' κ hg a Set.univ, measure_univ]⟩⟩ +instance IsZeroOrMarkovKernel.comap (κ : Kernel α β) [IsZeroOrMarkovKernel κ] (hg : Measurable g) : + IsZeroOrMarkovKernel (comap κ g hg) := by + rcases eq_zero_or_isMarkovKernel κ with rfl | h + · simp only [comap_zero]; infer_instance + · have := IsMarkovKernel.comap κ hg; infer_instance + instance IsFiniteKernel.comap (κ : Kernel α β) [IsFiniteKernel κ] (hg : Measurable g) : IsFiniteKernel (comap κ g hg) := by refine ⟨⟨IsFiniteKernel.bound κ, IsFiniteKernel.bound_lt_top κ, fun a => ?_⟩⟩ @@ -802,6 +808,12 @@ instance IsMarkovKernel.prodMkLeft (κ : Kernel α β) [IsMarkovKernel κ] : instance IsMarkovKernel.prodMkRight (κ : Kernel α β) [IsMarkovKernel κ] : IsMarkovKernel (prodMkRight γ κ) := by rw [Kernel.prodMkRight]; infer_instance +instance IsZeroOrMarkovKernel.prodMkLeft (κ : Kernel α β) [IsZeroOrMarkovKernel κ] : + IsZeroOrMarkovKernel (prodMkLeft γ κ) := by rw [Kernel.prodMkLeft]; infer_instance + +instance IsZeroOrMarkovKernel.prodMkRight (κ : Kernel α β) [IsZeroOrMarkovKernel κ] : + IsZeroOrMarkovKernel (prodMkRight γ κ) := by rw [Kernel.prodMkRight]; infer_instance + instance IsFiniteKernel.prodMkLeft (κ : Kernel α β) [IsFiniteKernel κ] : IsFiniteKernel (prodMkLeft γ κ) := by rw [Kernel.prodMkLeft]; infer_instance diff --git a/Mathlib/Probability/Kernel/Composition/ParallelComp.lean b/Mathlib/Probability/Kernel/Composition/ParallelComp.lean new file mode 100644 index 0000000000000..ab8450cc02791 --- /dev/null +++ b/Mathlib/Probability/Kernel/Composition/ParallelComp.lean @@ -0,0 +1,118 @@ +/- +Copyright (c) 2024 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne, Lorenzo Luccioli +-/ +import Mathlib.Probability.Kernel.Composition.Basic + +/-! + +# Parallel composition of kernels + +Two kernels `κ : Kernel α β` and `η : Kernel γ δ` can be applied in parallel to give a kernel +`κ ∥ₖ η` from `α × γ` to `β × δ`: `(κ ∥ₖ η) (a, c) = (κ a).prod (η c)`. + +## Main definitions + +* `parallelComp (κ : Kernel α β) (η : Kernel γ δ) : Kernel (α × γ) (β × δ)`: parallel composition + of two s-finite kernels. We define a notation `κ ∥ₖ η = parallelComp κ η`. + `∫⁻ bd, g bd ∂(κ ∥ₖ η) ac = ∫⁻ b, ∫⁻ d, g (b, d) ∂η ac.2 ∂κ ac.1` + +## Main statements + +* `parallelComp_comp_copy`: `(κ ∥ₖ η) ∘ₖ (copy α) = κ ×ₖ η` +* `deterministic_comp_copy`: for a deterministic kernel, copying then applying the kernel to + the two copies is the same as first applying the kernel then copying. That is, if `κ` is + a deterministic kernel, `(κ ∥ₖ κ) ∘ₖ copy α = copy β ∘ₖ κ`. + +## Notations + +* `κ ∥ₖ η = ProbabilityTheory.Kernel.parallelComp κ η` + +## Implementation notes + +Our formalization of kernels is centered around the composition-product: the product and then the +parallel composition are defined as special cases of the composition-product. +We could have alternatively used the building blocks of kernels seen as a Markov category: +composition, parallel composition (or tensor product) and the deterministic kernels `id`, `copy`, +`swap` and `discard`. The product and composition-product could then be built from these. + +-/ + +open MeasureTheory + +open scoped ENNReal + +namespace ProbabilityTheory.Kernel + +variable {α β γ δ : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} + {mγ : MeasurableSpace γ} {mδ : MeasurableSpace δ} + +section ParallelComp + +/-- Parallel product of two kernels. -/ +noncomputable +def parallelComp (κ : Kernel α β) (η : Kernel γ δ) : Kernel (α × γ) (β × δ) := + (prodMkRight γ κ) ×ₖ (prodMkLeft α η) + +@[inherit_doc] +scoped[ProbabilityTheory] infixl:100 " ∥ₖ " => ProbabilityTheory.Kernel.parallelComp + +lemma parallelComp_apply (κ : Kernel α β) [IsSFiniteKernel κ] + (η : Kernel γ δ) [IsSFiniteKernel η] (x : α × γ) : + (κ ∥ₖ η) x = (κ x.1).prod (η x.2) := by + rw [parallelComp, prod_apply, prodMkRight_apply, prodMkLeft_apply] + +lemma lintegral_parallelComp (κ : Kernel α β) [IsSFiniteKernel κ] + (η : Kernel γ δ) [IsSFiniteKernel η] + (ac : α × γ) {g : β × δ → ℝ≥0∞} (hg : Measurable g) : + ∫⁻ bd, g bd ∂(κ ∥ₖ η) ac = ∫⁻ b, ∫⁻ d, g (b, d) ∂η ac.2 ∂κ ac.1 := by + rw [parallelComp, lintegral_prod _ _ _ hg] + simp + +instance (κ : Kernel α β) (η : Kernel γ δ) : IsSFiniteKernel (κ ∥ₖ η) := by + rw [parallelComp]; infer_instance + +instance (κ : Kernel α β) [IsFiniteKernel κ] (η : Kernel γ δ) [IsFiniteKernel η] : + IsFiniteKernel (κ ∥ₖ η) := by + rw [parallelComp]; infer_instance + +instance (κ : Kernel α β) [IsMarkovKernel κ] (η : Kernel γ δ) [IsMarkovKernel η] : + IsMarkovKernel (κ ∥ₖ η) := by + rw [parallelComp]; infer_instance + +instance (κ : Kernel α β) [IsZeroOrMarkovKernel κ] (η : Kernel γ δ) [IsZeroOrMarkovKernel η] : + IsZeroOrMarkovKernel (κ ∥ₖ η) := by + rw [parallelComp]; infer_instance + +lemma parallelComp_comp_copy (κ : Kernel α β) [IsSFiniteKernel κ] + (η : Kernel α γ) [IsSFiniteKernel η] : + (κ ∥ₖ η) ∘ₖ (copy α) = κ ×ₖ η := by + ext a s hs + simp_rw [prod_apply, comp_apply, copy_apply, Measure.bind_apply hs (Kernel.measurable _)] + rw [lintegral_dirac'] + swap; · exact Kernel.measurable_coe _ hs + rw [parallelComp_apply] + +lemma swap_parallelComp {κ : Kernel α β} [IsSFiniteKernel κ] + {η : Kernel γ δ} [IsSFiniteKernel η] : + (swap β δ) ∘ₖ (κ ∥ₖ η) = (η ∥ₖ κ) ∘ₖ (swap α γ) := by + rw [parallelComp, swap_prod, parallelComp] + ext ac s hs + rw [comp_apply, swap_apply, Measure.bind_apply hs (Kernel.measurable _), + lintegral_dirac' _ (Kernel.measurable_coe _ hs), prod_apply, prod_apply, prodMkLeft_apply, + prodMkLeft_apply, prodMkRight_apply, prodMkRight_apply] + rfl + +/-- For a deterministic kernel, copying then applying the kernel to the two copies is the same +as first applying the kernel then copying. -/ +lemma deterministic_comp_copy {f : α → β} (hf : Measurable f) : + (Kernel.deterministic f hf ∥ₖ Kernel.deterministic f hf) ∘ₖ Kernel.copy α + = Kernel.copy β ∘ₖ Kernel.deterministic f hf := by + rw [Kernel.parallelComp_comp_copy, Kernel.deterministic_prod_deterministic, + Kernel.copy, Kernel.deterministic_comp_deterministic] + rfl + +end ParallelComp + +end ProbabilityTheory.Kernel From eaab7c8d11d0677c98095f371ef3dd56c32640a8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Wed, 11 Dec 2024 17:41:41 +0000 Subject: [PATCH 667/829] feat(Probability): absolute continuity of Measure.compProd (#19629) Also rename some lemmas to use dot notation. From the TestingLowerBounds project. Co-authored-by: Lorenzo Luccioli Co-authored-by: RemyDegenne --- .../MeasureTheory/Decomposition/Lebesgue.lean | 12 ++ .../Measure/MutuallySingular.lean | 21 +++ .../Kernel/Composition/MeasureCompProd.lean | 122 +++++++++++++++++- 3 files changed, 150 insertions(+), 5 deletions(-) diff --git a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean index e0b0eb5f2edfb..86fb861a6a9cb 100644 --- a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean @@ -229,6 +229,18 @@ lemma absolutelyContinuous_withDensity_rnDeriv [HaveLebesgueDecomposition ν μ] · exact measure_mono_null Set.inter_subset_left hνs · exact measure_mono_null Set.inter_subset_right ht2 +lemma AbsolutelyContinuous.withDensity_rnDeriv {ξ : Measure α} [μ.HaveLebesgueDecomposition ν] + (hξμ : ξ ≪ μ) (hξν : ξ ≪ ν) : + ξ ≪ ν.withDensity (μ.rnDeriv ν) := by + conv_rhs at hξμ => rw [μ.haveLebesgueDecomposition_add ν, add_comm] + refine absolutelyContinuous_of_add_of_mutuallySingular hξμ ?_ + exact MutuallySingular.mono_ac (mutuallySingular_singularPart μ ν).symm hξν .rfl + +lemma absolutelyContinuous_withDensity_rnDeriv_swap [ν.HaveLebesgueDecomposition μ] : + ν.withDensity (μ.rnDeriv ν) ≪ μ.withDensity (ν.rnDeriv μ) := + (withDensity_absolutelyContinuous ν (μ.rnDeriv ν)).withDensity_rnDeriv + (absolutelyContinuous_of_le (withDensity_rnDeriv_le _ _)) + lemma singularPart_eq_zero_of_ac (h : μ ≪ ν) : μ.singularPart ν = 0 := by rw [← MutuallySingular.self_iff] exact MutuallySingular.mono_ac (mutuallySingular_singularPart _ _) diff --git a/Mathlib/MeasureTheory/Measure/MutuallySingular.lean b/Mathlib/MeasureTheory/Measure/MutuallySingular.lean index e23107b1af78d..92ae8dd6e234a 100644 --- a/Mathlib/MeasureTheory/Measure/MutuallySingular.lean +++ b/Mathlib/MeasureTheory/Measure/MutuallySingular.lean @@ -89,6 +89,10 @@ theorem mono_ac (h : μ₁ ⟂ₘ ν₁) (hμ : μ₂ ≪ μ₁) (hν : ν₂ let ⟨s, hs, h₁, h₂⟩ := h ⟨s, hs, hμ h₁, hν h₂⟩ +lemma congr_ac (hμμ₂ : μ ≪ μ₂) (hμ₂μ : μ₂ ≪ μ) (hνν₂ : ν ≪ ν₂) (hν₂ν : ν₂ ≪ ν) : + μ ⟂ₘ ν ↔ μ₂ ⟂ₘ ν₂ := + ⟨fun h ↦ h.mono_ac hμ₂μ hν₂ν, fun h ↦ h.mono_ac hμμ₂ hνν₂⟩ + theorem mono (h : μ₁ ⟂ₘ ν₁) (hμ : μ₂ ≤ μ₁) (hν : ν₂ ≤ ν₁) : μ₂ ⟂ₘ ν₂ := h.mono_ac hμ.absolutelyContinuous hν.absolutelyContinuous @@ -146,6 +150,23 @@ lemma eq_zero_of_absolutelyContinuous_of_mutuallySingular {μ ν : Measure α} rw [← Measure.MutuallySingular.self_iff] exact h_ms.mono_ac Measure.AbsolutelyContinuous.rfl h_ac +lemma absolutelyContinuous_of_add_of_mutuallySingular {ν₁ ν₂ : Measure α} + (h : μ ≪ ν₁ + ν₂) (h_ms : μ ⟂ₘ ν₂) : μ ≪ ν₁ := by + refine AbsolutelyContinuous.mk fun s hs hs_zero ↦ ?_ + let t := h_ms.nullSet + have ht : MeasurableSet t := h_ms.measurableSet_nullSet + have htμ : μ t = 0 := h_ms.measure_nullSet + have htν₂ : ν₂ tᶜ = 0 := h_ms.measure_compl_nullSet + have : μ s = μ (s ∩ tᶜ) := by + conv_lhs => rw [← inter_union_compl s t] + rw [measure_union, measure_inter_null_of_null_right _ htμ, zero_add] + · exact (disjoint_compl_right.inter_right' _ ).inter_left' _ + · exact hs.inter ht.compl + rw [this] + refine h ?_ + simp only [Measure.coe_add, Pi.add_apply, add_eq_zero] + exact ⟨measure_inter_null_of_null_left _ hs_zero, measure_inter_null_of_null_right _ htν₂⟩ + lemma _root_.MeasurableEmbedding.mutuallySingular_map {β : Type*} {_ : MeasurableSpace β} {f : α → β} (hf : MeasurableEmbedding f) (hμν : μ ⟂ₘ ν) : μ.map f ⟂ₘ ν.map f := by diff --git a/Mathlib/Probability/Kernel/Composition/MeasureCompProd.lean b/Mathlib/Probability/Kernel/Composition/MeasureCompProd.lean index e15b968b5a1ba..b57ca2ef11459 100644 --- a/Mathlib/Probability/Kernel/Composition/MeasureCompProd.lean +++ b/Mathlib/Probability/Kernel/Composition/MeasureCompProd.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ +import Mathlib.MeasureTheory.Decomposition.Lebesgue import Mathlib.Probability.Kernel.Composition.IntegralCompProd /-! @@ -200,7 +201,7 @@ instance [IsProbabilityMeasure μ] [IsMarkovKernel κ] : IsProbabilityMeasure ( section AbsolutelyContinuous -lemma absolutelyContinuous_compProd_left [SFinite ν] (hμν : μ ≪ ν) (κ : Kernel α β) : +lemma AbsolutelyContinuous.compProd_left [SFinite ν] (hμν : μ ≪ ν) (κ : Kernel α β) : μ ⊗ₘ κ ≪ ν ⊗ₘ κ := by by_cases hκ : IsSFiniteKernel κ · have : SFinite μ := sFinite_of_absolutelyContinuous hμν @@ -210,7 +211,10 @@ lemma absolutelyContinuous_compProd_left [SFinite ν] (hμν : μ ≪ ν) (κ : exact hμν.ae_eq hs_zero · simp [compProd_of_not_isSFiniteKernel _ _ hκ] -lemma absolutelyContinuous_compProd_right [SFinite μ] [IsSFiniteKernel η] +@[deprecated (since := "2024-12-11")] +alias absolutelyContinuous_compProd_left := AbsolutelyContinuous.compProd_left + +lemma AbsolutelyContinuous.compProd_right [SFinite μ] [IsSFiniteKernel η] (hκη : ∀ᵐ a ∂μ, κ a ≪ η a) : μ ⊗ₘ κ ≪ μ ⊗ₘ η := by by_cases hκ : IsSFiniteKernel κ @@ -220,12 +224,17 @@ lemma absolutelyContinuous_compProd_right [SFinite μ] [IsSFiniteKernel η] filter_upwards [hs_zero, hκη] with a ha_zero ha_ac using ha_ac ha_zero · simp [compProd_of_not_isSFiniteKernel _ _ hκ] -lemma absolutelyContinuous_compProd [SFinite ν] [IsSFiniteKernel η] +@[deprecated (since := "2024-12-11")] +alias absolutelyContinuous_compProd_right := AbsolutelyContinuous.compProd_right + +lemma AbsolutelyContinuous.compProd [SFinite ν] [IsSFiniteKernel η] (hμν : μ ≪ ν) (hκη : ∀ᵐ a ∂μ, κ a ≪ η a) : μ ⊗ₘ κ ≪ ν ⊗ₘ η := have : SFinite μ := sFinite_of_absolutelyContinuous hμν - (Measure.absolutelyContinuous_compProd_right hκη).trans - (Measure.absolutelyContinuous_compProd_left hμν _) + (Measure.AbsolutelyContinuous.compProd_right hκη).trans (hμν.compProd_left _) + +@[deprecated (since := "2024-12-11")] +alias absolutelyContinuous_compProd := AbsolutelyContinuous.compProd lemma absolutelyContinuous_of_compProd [SFinite μ] [IsSFiniteKernel κ] [h_zero : ∀ a, NeZero (κ a)] (h : μ ⊗ₘ κ ≪ ν ⊗ₘ η) : @@ -248,6 +257,109 @@ lemma absolutelyContinuous_of_compProd [SFinite μ] [IsSFiniteKernel κ] [h_zero simp only [Measure.measure_univ_eq_zero] exact (h_zero a).out +lemma absolutelyContinuous_compProd_left_iff [SFinite μ] [SFinite ν] + [IsFiniteKernel κ] [∀ a, NeZero (κ a)] : + μ ⊗ₘ κ ≪ ν ⊗ₘ κ ↔ μ ≪ ν := + ⟨absolutelyContinuous_of_compProd, fun h ↦ h.compProd_left κ⟩ + +lemma AbsolutelyContinuous.compProd_of_compProd [SFinite ν] [IsSFiniteKernel η] + (hμν : μ ≪ ν) (hκη : μ ⊗ₘ κ ≪ μ ⊗ₘ η) : + μ ⊗ₘ κ ≪ ν ⊗ₘ η := by + by_cases hμ : SFinite μ + swap; · rw [compProd_of_not_sfinite _ _ hμ]; simp + refine AbsolutelyContinuous.mk fun s hs hs_zero ↦ ?_ + suffices (μ ⊗ₘ η) s = 0 from hκη this + rw [measure_zero_iff_ae_nmem, ae_compProd_iff hs.compl] at hs_zero ⊢ + exact hμν.ae_le hs_zero + end AbsolutelyContinuous +section MutuallySingular + +lemma MutuallySingular.compProd_of_left (hμν : μ ⟂ₘ ν) (κ η : Kernel α β) : + μ ⊗ₘ κ ⟂ₘ ν ⊗ₘ η := by + by_cases hμ : SFinite μ + swap; · rw [compProd_of_not_sfinite _ _ hμ]; simp + by_cases hν : SFinite ν + swap; · rw [compProd_of_not_sfinite _ _ hν]; simp + by_cases hκ : IsSFiniteKernel κ + swap; · rw [compProd_of_not_isSFiniteKernel _ _ hκ]; simp + by_cases hη : IsSFiniteKernel η + swap; · rw [compProd_of_not_isSFiniteKernel _ _ hη]; simp + refine ⟨hμν.nullSet ×ˢ univ, hμν.measurableSet_nullSet.prod .univ, ?_⟩ + rw [compProd_apply_prod hμν.measurableSet_nullSet .univ, compl_prod_eq_union] + simp only [MutuallySingular.restrict_nullSet, lintegral_zero_measure, compl_univ, + prod_empty, union_empty, true_and] + rw [compProd_apply_prod hμν.measurableSet_nullSet.compl .univ] + simp + +lemma mutuallySingular_of_mutuallySingular_compProd {ξ : Measure α} + [SFinite μ] [SFinite ν] [IsSFiniteKernel κ] [IsSFiniteKernel η] + (h : μ ⊗ₘ κ ⟂ₘ ν ⊗ₘ η) (hμ : ξ ≪ μ) (hν : ξ ≪ ν) : + ∀ᵐ x ∂ξ, κ x ⟂ₘ η x := by + have hs : MeasurableSet h.nullSet := h.measurableSet_nullSet + have hμ_zero : (μ ⊗ₘ κ) h.nullSet = 0 := h.measure_nullSet + have hν_zero : (ν ⊗ₘ η) h.nullSetᶜ = 0 := h.measure_compl_nullSet + rw [compProd_apply, lintegral_eq_zero_iff'] at hμ_zero hν_zero + · filter_upwards [hμ hμ_zero, hν hν_zero] with x hxμ hxν + exact ⟨Prod.mk x ⁻¹' h.nullSet, measurable_prod_mk_left hs, ⟨hxμ, hxν⟩⟩ + · exact (Kernel.measurable_kernel_prod_mk_left hs.compl).aemeasurable + · exact (Kernel.measurable_kernel_prod_mk_left hs).aemeasurable + · exact hs.compl + · exact hs + +lemma mutuallySingular_compProd_left_iff [SFinite μ] [SigmaFinite ν] + [IsSFiniteKernel κ] [hκ : ∀ x, NeZero (κ x)] : + μ ⊗ₘ κ ⟂ₘ ν ⊗ₘ κ ↔ μ ⟂ₘ ν := by + refine ⟨fun h ↦ ?_, fun h ↦ h.compProd_of_left _ _⟩ + rw [← withDensity_rnDeriv_eq_zero] + have hh := mutuallySingular_of_mutuallySingular_compProd h ?_ ?_ + (ξ := ν.withDensity (μ.rnDeriv ν)) + rotate_left + · exact absolutelyContinuous_of_le (μ.withDensity_rnDeriv_le ν) + · exact withDensity_absolutelyContinuous _ _ + simp_rw [MutuallySingular.self_iff, (hκ _).ne] at hh + exact ae_eq_bot.mp (Filter.eventually_false_iff_eq_bot.mp hh) + +lemma AbsolutelyContinuous.mutuallySingular_compProd_iff [SigmaFinite μ] [SigmaFinite ν] + (hμν : μ ≪ ν) : + μ ⊗ₘ κ ⟂ₘ ν ⊗ₘ η ↔ μ ⊗ₘ κ ⟂ₘ μ ⊗ₘ η := by + conv_lhs => rw [ν.haveLebesgueDecomposition_add μ] + rw [compProd_add_left, MutuallySingular.add_right_iff] + simp only [(mutuallySingular_singularPart ν μ).symm.compProd_of_left κ η, true_and] + refine ⟨fun h ↦ h.mono_ac .rfl ?_, fun h ↦ h.mono_ac .rfl ?_⟩ + · exact (absolutelyContinuous_withDensity_rnDeriv hμν).compProd_left _ + · exact (withDensity_absolutelyContinuous μ (ν.rnDeriv μ)).compProd_left _ + +lemma mutuallySingular_compProd_iff [SigmaFinite μ] [SigmaFinite ν] : + μ ⊗ₘ κ ⟂ₘ ν ⊗ₘ η ↔ ∀ ξ, SFinite ξ → ξ ≪ μ → ξ ≪ ν → ξ ⊗ₘ κ ⟂ₘ ξ ⊗ₘ η := by + conv_lhs => rw [μ.haveLebesgueDecomposition_add ν] + rw [compProd_add_left, MutuallySingular.add_left_iff] + simp only [(mutuallySingular_singularPart μ ν).compProd_of_left κ η, true_and] + rw [(withDensity_absolutelyContinuous ν (μ.rnDeriv ν)).mutuallySingular_compProd_iff] + refine ⟨fun h ξ hξ hξμ hξν ↦ ?_, fun h ↦ ?_⟩ + · exact h.mono_ac ((hξμ.withDensity_rnDeriv hξν).compProd_left _) + ((hξμ.withDensity_rnDeriv hξν).compProd_left _) + · refine h _ ?_ ?_ ?_ + · infer_instance + · exact absolutelyContinuous_of_le (withDensity_rnDeriv_le _ _) + · exact withDensity_absolutelyContinuous ν (μ.rnDeriv ν) + +end MutuallySingular + +lemma absolutelyContinuous_compProd_of_compProd [SigmaFinite μ] [SigmaFinite ν] + (hκη : μ ⊗ₘ κ ≪ ν ⊗ₘ η) : + μ ⊗ₘ κ ≪ μ ⊗ₘ η := by + rw [ν.haveLebesgueDecomposition_add μ, compProd_add_left, add_comm] at hκη + have h := absolutelyContinuous_of_add_of_mutuallySingular hκη + ((mutuallySingular_singularPart _ _).symm.compProd_of_left _ _) + refine h.trans (AbsolutelyContinuous.compProd_left ?_ _) + exact withDensity_absolutelyContinuous _ _ + +lemma absolutelyContinuous_compProd_iff + [SigmaFinite μ] [SigmaFinite ν] [IsSFiniteKernel κ] [IsSFiniteKernel η] [∀ x, NeZero (κ x)] : + μ ⊗ₘ κ ≪ ν ⊗ₘ η ↔ μ ≪ ν ∧ μ ⊗ₘ κ ≪ μ ⊗ₘ η := + ⟨fun h ↦ ⟨absolutelyContinuous_of_compProd h, absolutelyContinuous_compProd_of_compProd h⟩, + fun h ↦ h.1.compProd_of_compProd h.2⟩ + end MeasureTheory.Measure From ca6f586b306dbe71c3465b5bead67c64097e1570 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 11 Dec 2024 17:56:48 +0000 Subject: [PATCH 668/829] chore(*): fix `initialize_simps_projections` commands (#19422) With old configs, `@[simps]` generated both `_toFun` and `_apply` lemmas. --- Mathlib/Algebra/Order/AbsoluteValue.lean | 2 +- .../Analysis/Calculus/ContDiff/FaaDiBruno.lean | 13 +++++++------ .../Normed/Operator/LinearIsometry.lean | 5 ++--- .../AffineSpace/ContinuousAffineEquiv.lean | 2 +- Mathlib/MeasureTheory/Measure/Haar/Unique.lean | 4 ++-- Mathlib/MeasureTheory/Measure/Portmanteau.lean | 2 +- Mathlib/NumberTheory/Padics/MahlerBasis.lean | 2 +- Mathlib/Topology/Algebra/Algebra.lean | 2 +- Mathlib/Topology/Algebra/Module/Basic.lean | 2 +- .../Topology/ContinuousMap/Bounded/Basic.lean | 17 +++++++++-------- Mathlib/Topology/ContinuousMap/Ideals.lean | 2 +- Mathlib/Topology/Homotopy/Basic.lean | 3 +-- Mathlib/Topology/MetricSpace/Isometry.lean | 2 +- 13 files changed, 29 insertions(+), 29 deletions(-) diff --git a/Mathlib/Algebra/Order/AbsoluteValue.lean b/Mathlib/Algebra/Order/AbsoluteValue.lean index 8f31d949787d9..657398ad3fd64 100644 --- a/Mathlib/Algebra/Order/AbsoluteValue.lean +++ b/Mathlib/Algebra/Order/AbsoluteValue.lean @@ -75,7 +75,7 @@ theorem ext ⦃f g : AbsoluteValue R S⦄ : (∀ x, f x = g x) → f = g := def Simps.apply (f : AbsoluteValue R S) : R → S := f -initialize_simps_projections AbsoluteValue (toMulHom_toFun → apply) +initialize_simps_projections AbsoluteValue (toFun → apply) @[simp] theorem coe_toMulHom : ⇑abv.toMulHom = abv := diff --git a/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean b/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean index f9ce84b67909d..f8b0483d56f73 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean @@ -895,12 +895,13 @@ private lemma faaDiBruno_aux2 {m : ℕ} (q : FormalMultilinearSeries 𝕜 F G) ((c.compAlongOrderedFinpartitionL 𝕜 E F G (q c.length)).toContinuousLinearMap (fun i ↦ p (c.partSize i)) i).comp (p (c.partSize i + 1)).curryLeft := by ext e v - simp only [Nat.succ_eq_add_one, OrderedFinpartition.extend, extendMiddle, - ContinuousMultilinearMap.curryLeft_apply, - FormalMultilinearSeries.compAlongOrderedFinpartition_apply, ContinuousLinearMap.coe_comp', - comp_apply, ContinuousMultilinearMap.toContinuousLinearMap_toFun, - compAlongOrderedFinpartitionL_apply, compAlongOrderFinpartition_apply, - applyOrderedFinpartition_apply] + simp? [OrderedFinpartition.extend, extendMiddle, applyOrderedFinpartition_apply] says + simp only [Nat.succ_eq_add_one, OrderedFinpartition.extend, extendMiddle, + ContinuousMultilinearMap.curryLeft_apply, + FormalMultilinearSeries.compAlongOrderedFinpartition_apply, applyOrderedFinpartition_apply, + ContinuousLinearMap.coe_comp', comp_apply, + ContinuousMultilinearMap.toContinuousLinearMap_apply, compAlongOrderedFinpartitionL_apply, + compAlongOrderFinpartition_apply] congr ext j rcases eq_or_ne j i with rfl | hij diff --git a/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean index 26c2ba5bcceec..7628196bc3d17 100644 --- a/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean +++ b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean @@ -162,7 +162,7 @@ def Simps.apply (σ₁₂ : R →+* R₂) (E E₂ : Type*) [SeminormedAddCommGro [SeminormedAddCommGroup E₂] [Module R E] [Module R₂ E₂] (h : E →ₛₗᵢ[σ₁₂] E₂) : E → E₂ := h -initialize_simps_projections LinearIsometry (toLinearMap_toFun → apply) +initialize_simps_projections LinearIsometry (toFun → apply) @[ext] theorem ext {f g : E →ₛₗᵢ[σ₁₂] E₂} (h : ∀ x, f x = g x) : f = g := @@ -668,8 +668,7 @@ def Simps.symm_apply (σ₁₂ : R →+* R₂) {σ₂₁ : R₂ →+* R} [RingHo [Module R E] [Module R₂ E₂] (h : E ≃ₛₗᵢ[σ₁₂] E₂) : E₂ → E := h.symm -initialize_simps_projections LinearIsometryEquiv (toLinearEquiv_toFun → apply, - toLinearEquiv_invFun → symm_apply) +initialize_simps_projections LinearIsometryEquiv (toFun → apply, invFun → symm_apply) /-- Composition of `LinearIsometryEquiv`s as a `LinearIsometryEquiv`. -/ def trans (e' : E₂ ≃ₛₗᵢ[σ₂₃] E₃) : E ≃ₛₗᵢ[σ₁₃] E₃ := diff --git a/Mathlib/LinearAlgebra/AffineSpace/ContinuousAffineEquiv.lean b/Mathlib/LinearAlgebra/AffineSpace/ContinuousAffineEquiv.lean index df744660c3393..5631cbfd7acca 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/ContinuousAffineEquiv.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/ContinuousAffineEquiv.lean @@ -105,7 +105,7 @@ def Simps.apply (e : P₁ ≃ᵃL[k] P₂) : P₁ → P₂ := def Simps.coe (e : P₁ ≃ᵃL[k] P₂) : P₁ ≃ᵃ[k] P₂ := e -initialize_simps_projections ContinuousLinearMap (toAffineEquiv_toFun → apply, toAffineEquiv → coe) +initialize_simps_projections ContinuousLinearMap (toFun → apply, toAffineEquiv → coe) @[ext] theorem ext {e e' : P₁ ≃ᵃL[k] P₂} (h : ∀ x, e x = e' x) : e = e' := diff --git a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean index 377e30300f09c..e9d2e87202f9e 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean @@ -483,13 +483,13 @@ lemma measure_preimage_isMulLeftInvariant_eq_smul_of_hasCompactSupport · filter_upwards with x have T := tendsto_pi_nhds.1 (thickenedIndicator_tendsto_indicator_closure (fun n ↦ (u_mem n).1) u_lim ({1} : Set ℝ)) (f x) - simp only [thickenedIndicator_toFun, closure_singleton] at T + simp only [thickenedIndicator_apply, closure_singleton] at T convert NNReal.tendsto_coe.2 T simp have M n : ∫ (x : G), v n (f x) ∂μ' = ∫ (x : G), v n (f x) ∂(haarScalarFactor μ' μ • μ) := by apply integral_isMulLeftInvariant_eq_smul_of_hasCompactSupport μ' μ (vf_cont n) apply h'f.comp_left - simp only [v, thickenedIndicator_toFun, NNReal.coe_eq_zero] + simp only [v, thickenedIndicator_apply, NNReal.coe_eq_zero] rw [thickenedIndicatorAux_zero (u_mem n).1] · simp only [ENNReal.zero_toNNReal] · simpa using (u_mem n).2.le diff --git a/Mathlib/MeasureTheory/Measure/Portmanteau.lean b/Mathlib/MeasureTheory/Measure/Portmanteau.lean index 6cb6af9e29fd8..55aff1558ed0e 100644 --- a/Mathlib/MeasureTheory/Measure/Portmanteau.lean +++ b/Mathlib/MeasureTheory/Measure/Portmanteau.lean @@ -508,7 +508,7 @@ lemma integral_le_liminf_integral_of_forall_isOpen_measure_le_liminf_measure apply le_trans hi convert obs i with x have aux := ENNReal.ofReal_eq_coe_nnreal (f_nn x) - simp only [ContinuousMap.toFun_eq_coe, BoundedContinuousFunction.coe_to_continuous_fun] at aux + simp only [ContinuousMap.toFun_eq_coe, BoundedContinuousFunction.coe_toContinuousMap] at aux rw [aux] congr exact (Real.norm_of_nonneg (f_nn x)).symm diff --git a/Mathlib/NumberTheory/Padics/MahlerBasis.lean b/Mathlib/NumberTheory/Padics/MahlerBasis.lean index 55ebe61b44164..c96931b3ec854 100644 --- a/Mathlib/NumberTheory/Padics/MahlerBasis.lean +++ b/Mathlib/NumberTheory/Padics/MahlerBasis.lean @@ -367,7 +367,7 @@ noncomputable def mahlerEquiv : C(ℤ_[p], E) ≃ₗᵢ[ℚ_[p]] C₀(ℕ, E) wh · rw [← (hasSum_mahler f).tsum_eq] refine (norm_tsum_le _).trans (ciSup_le fun n ↦ ?_) refine le_trans (le_of_eq ?_) (BoundedContinuousFunction.norm_coe_le_norm _ n) - simp only [ZeroAtInftyContinuousMap.toBCF_toFun, ZeroAtInftyContinuousMap.coe_mk, + simp only [ZeroAtInftyContinuousMap.toBCF_apply, ZeroAtInftyContinuousMap.coe_mk, norm_mahlerTerm, (hasSum_mahler f).tsum_eq] lemma mahlerEquiv_apply (f : C(ℤ_[p], E)) : mahlerEquiv E f = fun n ↦ Δ_[1]^[n] f 0 := rfl diff --git a/Mathlib/Topology/Algebra/Algebra.lean b/Mathlib/Topology/Algebra/Algebra.lean index ceec483fcb959..9e7f97fa9ee48 100644 --- a/Mathlib/Topology/Algebra/Algebra.lean +++ b/Mathlib/Topology/Algebra/Algebra.lean @@ -157,7 +157,7 @@ def Simps.apply (h : A →A[R] B) : A → B := h /-- See Note [custom simps projection]. -/ def Simps.coe (h : A →A[R] B) : A →ₐ[R] B := h -initialize_simps_projections ContinuousAlgHom (toAlgHom_toFun → apply, toAlgHom → coe) +initialize_simps_projections ContinuousAlgHom (toFun → apply, toAlgHom → coe) @[ext] theorem ext {f g : A →A[R] B} (h : ∀ x, f x = g x) : f = g := DFunLike.ext f g h diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index 4dbcb4ad07c9d..9e0b39d1f56d0 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -408,7 +408,7 @@ def Simps.apply (h : M₁ →SL[σ₁₂] M₂) : M₁ → M₂ := def Simps.coe (h : M₁ →SL[σ₁₂] M₂) : M₁ →ₛₗ[σ₁₂] M₂ := h -initialize_simps_projections ContinuousLinearMap (toLinearMap_toFun → apply, toLinearMap → coe) +initialize_simps_projections ContinuousLinearMap (toFun → apply, toLinearMap → coe) @[ext] theorem ext {f g : M₁ →SL[σ₁₂] M₂} (h : ∀ x, f x = g x) : f = g := diff --git a/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean b/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean index b378f1cd1420d..ab3605b44b21f 100644 --- a/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean +++ b/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean @@ -83,13 +83,15 @@ instance instCoeTC [FunLike F α β] [BoundedContinuousMapClass F α β] : CoeTC map_bounded' := map_bounded f }⟩ @[simp] -theorem coe_to_continuous_fun (f : α →ᵇ β) : (f.toContinuousMap : α → β) = f := rfl +theorem coe_toContinuousMap (f : α →ᵇ β) : (f.toContinuousMap : α → β) = f := rfl + +@[deprecated (since := "2024-11-23")] alias coe_to_continuous_fun := coe_toContinuousMap /-- See Note [custom simps projection]. We need to specify this projection explicitly in this case, because it is a composition of multiple projections. -/ def Simps.apply (h : α →ᵇ β) : α → β := h -initialize_simps_projections BoundedContinuousFunction (toContinuousMap_toFun → apply) +initialize_simps_projections BoundedContinuousFunction (toFun → apply) protected theorem bounded (f : α →ᵇ β) : ∃ C, ∀ x y : α, dist (f x) (f y) ≤ C := f.map_bounded' @@ -1420,8 +1422,7 @@ instance instNormedLatticeAddCommGroup : NormedLatticeAddCommGroup (α →ᵇ β { instSeminormedAddCommGroup with add_le_add_left := by intro f g h₁ h t - simp only [coe_to_continuous_fun, Pi.add_apply, add_le_add_iff_left, coe_add, - ContinuousMap.toFun_eq_coe] + simp only [coe_toContinuousMap, Pi.add_apply, add_le_add_iff_left, coe_add] exact h₁ _ solid := by intro f g h @@ -1480,15 +1481,15 @@ variable {α : Type*} [TopologicalSpace α] lemma add_norm_nonneg (f : α →ᵇ ℝ) : 0 ≤ f + const _ ‖f‖ := by intro x - simp only [ContinuousMap.toFun_eq_coe, coe_to_continuous_fun, coe_zero, Pi.zero_apply, coe_add, - const_toFun, Pi.add_apply] + simp only [ContinuousMap.toFun_eq_coe, coe_toContinuousMap, coe_zero, Pi.zero_apply, coe_add, + const_apply, Pi.add_apply] linarith [(abs_le.mp (norm_coe_le_norm f x)).1] lemma norm_sub_nonneg (f : α →ᵇ ℝ) : 0 ≤ const _ ‖f‖ - f := by intro x - simp only [ContinuousMap.toFun_eq_coe, coe_to_continuous_fun, coe_zero, Pi.zero_apply, coe_sub, - const_toFun, Pi.sub_apply, sub_nonneg] + simp only [ContinuousMap.toFun_eq_coe, coe_toContinuousMap, coe_zero, Pi.zero_apply, coe_sub, + const_apply, Pi.sub_apply, sub_nonneg] linarith [(abs_le.mp (norm_coe_le_norm f x)).2] end diff --git a/Mathlib/Topology/ContinuousMap/Ideals.lean b/Mathlib/Topology/ContinuousMap/Ideals.lean index 44324eab0595f..86b7ba981bb03 100644 --- a/Mathlib/Topology/ContinuousMap/Ideals.lean +++ b/Mathlib/Topology/ContinuousMap/Ideals.lean @@ -265,7 +265,7 @@ theorem idealOfSet_ofIdeal_eq_closure (I : Ideal C(X, 𝕜)) : pow_pos (norm_pos_iff.mpr hx.1) 2⟩⟩ convert I.mul_mem_left (star g) hI ext - simp only [comp_apply, ContinuousMap.coe_coe, coe_mk, algebraMapCLM_toFun, map_pow, + simp only [comp_apply, ContinuousMap.coe_coe, coe_mk, algebraMapCLM_apply, map_pow, mul_apply, star_apply, star_def] simp only [normSq_eq_def', RCLike.conj_mul, ofReal_pow] rfl diff --git a/Mathlib/Topology/Homotopy/Basic.lean b/Mathlib/Topology/Homotopy/Basic.lean index 2f4f42e68d200..20fa89d5458df 100644 --- a/Mathlib/Topology/Homotopy/Basic.lean +++ b/Mathlib/Topology/Homotopy/Basic.lean @@ -399,8 +399,7 @@ theorem ext {F G : HomotopyWith f₀ f₁ P} (h : ∀ x, F x = G x) : F = G := D because it is a composition of multiple projections. -/ def Simps.apply (F : HomotopyWith f₀ f₁ P) : I × X → Y := F -initialize_simps_projections HomotopyWith (toHomotopy_toContinuousMap_toFun → apply, - -toHomotopy_toContinuousMap) +initialize_simps_projections HomotopyWith (toFun → apply, -toHomotopy_toContinuousMap) @[continuity] protected theorem continuous (F : HomotopyWith f₀ f₁ P) : Continuous F := diff --git a/Mathlib/Topology/MetricSpace/Isometry.lean b/Mathlib/Topology/MetricSpace/Isometry.lean index a987f5c09f284..e145a00ce470a 100644 --- a/Mathlib/Topology/MetricSpace/Isometry.lean +++ b/Mathlib/Topology/MetricSpace/Isometry.lean @@ -365,7 +365,7 @@ def Simps.apply (h : α ≃ᵢ β) : α → β := h def Simps.symm_apply (h : α ≃ᵢ β) : β → α := h.symm -initialize_simps_projections IsometryEquiv (toEquiv_toFun → apply, toEquiv_invFun → symm_apply) +initialize_simps_projections IsometryEquiv (toFun → apply, invFun → symm_apply) @[simp] theorem symm_symm (h : α ≃ᵢ β) : h.symm.symm = h := rfl From 7fe7d99f2a27e141d574f45cb8b541f695c71511 Mon Sep 17 00:00:00 2001 From: Seewoo Lee Date: Wed, 11 Dec 2024 18:39:58 +0000 Subject: [PATCH 669/829] chore(Radical): minor refactoring of lemmas (#19875) --- Mathlib/RingTheory/Radical.lean | 18 +++++++++--------- 1 file changed, 9 insertions(+), 9 deletions(-) diff --git a/Mathlib/RingTheory/Radical.lean b/Mathlib/RingTheory/Radical.lean index a4558ba7c2ec9..cea98e424f69b 100644 --- a/Mathlib/RingTheory/Radical.lean +++ b/Mathlib/RingTheory/Radical.lean @@ -83,14 +83,14 @@ theorem radical_eq_of_associated {a b : M} (h : Associated a b) : radical a = ra unfold radical rw [h.primeFactors_eq] -theorem radical_unit_eq_one {a : M} (h : IsUnit a) : radical a = 1 := +theorem radical_of_isUnit {a : M} (h : IsUnit a) : radical a = 1 := (radical_eq_of_associated (associated_one_iff_isUnit.mpr h)).trans radical_one_eq -theorem radical_unit_mul {u : Mˣ} {a : M} : radical ((↑u : M) * a) = radical a := - radical_eq_of_associated (associated_unit_mul_left _ _ u.isUnit) +theorem radical_mul_of_isUnit_left {a u : M} (h : IsUnit u) : radical (u * a) = radical a := + radical_eq_of_associated (associated_unit_mul_left _ _ h) -theorem radical_mul_unit {u : Mˣ} {a : M} : radical (a * (↑u : M)) = radical a := - radical_eq_of_associated (associated_mul_unit_left _ _ u.isUnit) +theorem radical_mul_of_isUnit_right {a u : M} (h : IsUnit u) : radical (a * u) = radical a := + radical_eq_of_associated (associated_mul_unit_left _ _ h) theorem primeFactors_pow (a : M) {n : ℕ} (hn : 0 < n) : primeFactors (a ^ n) = primeFactors a := by simp_rw [primeFactors] @@ -162,17 +162,17 @@ theorem mul_primeFactors_disjUnion {a b : R} (ha : a ≠ 0) (hb : b ≠ 0) @[simp] theorem radical_neg_one : radical (-1 : R) = 1 := - radical_unit_eq_one isUnit_one.neg + radical_of_isUnit isUnit_one.neg /-- Radical is multiplicative for coprime elements. -/ theorem radical_mul {a b : R} (hc : IsCoprime a b) : radical (a * b) = radical a * radical b := by by_cases ha : a = 0 · subst ha; rw [isCoprime_zero_left] at hc - simp only [zero_mul, radical_zero_eq, one_mul, radical_unit_eq_one hc] + simp only [zero_mul, radical_zero_eq, one_mul, radical_of_isUnit hc] by_cases hb : b = 0 · subst hb; rw [isCoprime_zero_right] at hc - simp only [mul_zero, radical_zero_eq, mul_one, radical_unit_eq_one hc] + simp only [mul_zero, radical_zero_eq, mul_one, radical_of_isUnit hc] simp_rw [radical] rw [mul_primeFactors_disjUnion ha hb hc] rw [Finset.prod_disjUnion (disjoint_primeFactors hc)] @@ -205,7 +205,7 @@ theorem divRadical_ne_zero {a : E} (ha : a ≠ 0) : divRadical a ≠ 0 := by exact right_ne_zero_of_mul ha theorem divRadical_isUnit {u : E} (hu : IsUnit u) : IsUnit (divRadical u) := by - rwa [divRadical, radical_unit_eq_one hu, EuclideanDomain.div_one] + rwa [divRadical, radical_of_isUnit hu, EuclideanDomain.div_one] theorem eq_divRadical {a x : E} (h : radical a * x = a) : x = divRadical a := by apply EuclideanDomain.eq_div_of_mul_eq_left (radical_ne_zero a) From 4411562e4fd2d02aa31e413bc8e948829aa3a62d Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Wed, 11 Dec 2024 19:01:06 +0000 Subject: [PATCH 670/829] feat(Condensed): light condensed modules satisfy countable AB4* (#18497) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Given sequences `M N : ℕ → C` of objects with morphisms `f n : M n ⟶ N n` for all `n`, we exhibit `∏ M` as the limit of the tower ``` ⋯ → ∏_{n < m + 1} M n × ∏_{n ≥ m + 1} N n → ∏_{n < m} M n × ∏_{n ≥ m} N n → ⋯ → ∏ N ``` Further, the transition maps in this tower are epimorphisms, in the case when each `f n` is an epimorphism and `C` has finite biproducts. We use this to conclude that light condensed modules over a ring R satisfy a countable version of Grothendieck's AB4* axiom. --- Mathlib.lean | 2 + .../Limits/Shapes/SequentialProduct.lean | 230 ++++++++++++++++++ Mathlib/Condensed/Light/AB.lean | 30 +++ Mathlib/Condensed/Light/Epi.lean | 39 ++- Mathlib/Condensed/Light/Limits.lean | 8 +- 5 files changed, 303 insertions(+), 6 deletions(-) create mode 100644 Mathlib/CategoryTheory/Limits/Shapes/SequentialProduct.lean create mode 100644 Mathlib/Condensed/Light/AB.lean diff --git a/Mathlib.lean b/Mathlib.lean index 673308d1c2f37..236d70218b0b1 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1898,6 +1898,7 @@ import Mathlib.CategoryTheory.Limits.Shapes.Pullback.PullbackCone import Mathlib.CategoryTheory.Limits.Shapes.Pullback.Square import Mathlib.CategoryTheory.Limits.Shapes.Reflexive import Mathlib.CategoryTheory.Limits.Shapes.RegularMono +import Mathlib.CategoryTheory.Limits.Shapes.SequentialProduct import Mathlib.CategoryTheory.Limits.Shapes.SingleObj import Mathlib.CategoryTheory.Limits.Shapes.SplitCoequalizer import Mathlib.CategoryTheory.Limits.Shapes.SplitEqualizer @@ -2312,6 +2313,7 @@ import Mathlib.Condensed.Epi import Mathlib.Condensed.Equivalence import Mathlib.Condensed.Explicit import Mathlib.Condensed.Functors +import Mathlib.Condensed.Light.AB import Mathlib.Condensed.Light.Basic import Mathlib.Condensed.Light.CartesianClosed import Mathlib.Condensed.Light.Epi diff --git a/Mathlib/CategoryTheory/Limits/Shapes/SequentialProduct.lean b/Mathlib/CategoryTheory/Limits/Shapes/SequentialProduct.lean new file mode 100644 index 0000000000000..ed1e65b9bebfe --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Shapes/SequentialProduct.lean @@ -0,0 +1,230 @@ +/- +Copyright (c) 2024 Dagur Asgeirsson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Dagur Asgeirsson +-/ +import Mathlib.Algebra.Order.Monoid.Canonical.Defs +import Mathlib.CategoryTheory.Functor.OfSequence +import Mathlib.CategoryTheory.Limits.Shapes.Biproducts +import Mathlib.CategoryTheory.Limits.Shapes.Countable +import Mathlib.CategoryTheory.Limits.Shapes.PiProd +import Mathlib.CategoryTheory.Limits.Shapes.RegularMono +import Mathlib.Order.Interval.Finset.Nat +/-! + +# ℕ-indexed products as sequential limits + +Given sequences `M N : ℕ → C` of objects with morphisms `f n : M n ⟶ N n` for all `n`, this file +exhibits `∏ M` as the limit of the tower + +``` +⋯ → ∏_{n < m + 1} M n × ∏_{n ≥ m + 1} N n → ∏_{n < m} M n × ∏_{n ≥ m} N n → ⋯ → ∏ N +``` + +Further, we prove that the transition maps in this tower are epimorphisms, in the case when each +`f n` is an epimorphism and `C` has finite biproducts. +-/ + +namespace CategoryTheory.Limits.SequentialProduct + +variable {C : Type*} {M N : ℕ → C} + +lemma functorObj_eq_pos {n m : ℕ} (h : m < n) : + (fun i ↦ if _ : i < n then M i else N i) m = M m := dif_pos h + +lemma functorObj_eq_neg {n m : ℕ} (h : ¬(m < n)) : + (fun i ↦ if _ : i < n then M i else N i) m = N m := dif_neg h + +variable [Category C] (f : ∀ n, M n ⟶ N n) [HasProductsOfShape ℕ C] + +variable (M N) in +/-- The product of the `m` first objects of `M` and the rest of the rest of `N` -/ +noncomputable def functorObj : ℕ → C := + fun n ↦ ∏ᶜ (fun m ↦ if _ : m < n then M m else N m) + +/-- The projection map from `functorObj M N n` to `M m`, when `m < n` -/ +noncomputable def functorObjProj_pos (n m : ℕ) (h : m < n) : + functorObj M N n ⟶ M m := + Pi.π (fun m ↦ if _ : m < n then M m else N m) m ≫ eqToHom (functorObj_eq_pos (by omega)) + +/-- The projection map from `functorObj M N n` to `N m`, when `m ≥ n` -/ +noncomputable def functorObjProj_neg (n m : ℕ) (h : ¬(m < n)) : + functorObj M N n ⟶ N m := + Pi.π (fun m ↦ if _ : m < n then M m else N m) m ≫ eqToHom (functorObj_eq_neg (by omega)) + +/-- The transition maps in the sequential limit of products -/ +noncomputable def functorMap : ∀ n, + functorObj M N (n + 1) ⟶ functorObj M N n := by + intro n + refine Limits.Pi.map fun m ↦ if h : m < n then eqToHom ?_ else + if h' : m < n + 1 then eqToHom ?_ ≫ f m ≫ eqToHom ?_ else eqToHom ?_ + all_goals split_ifs; try rfl; try omega + +lemma functorMap_commSq_succ (n : ℕ) : + (Functor.ofOpSequence (functorMap f)).map (homOfLE (by omega : n ≤ n+1)).op ≫ Pi.π _ n ≫ + eqToHom (functorObj_eq_neg (by omega : ¬(n < n))) = + (Pi.π (fun i ↦ if _ : i < (n + 1) then M i else N i) n) ≫ + eqToHom (functorObj_eq_pos (by omega)) ≫ f n := by + simp [functorMap] + +lemma functorMap_commSq_aux {n m k : ℕ} (h : n ≤ m) (hh : ¬(k < m)) : + (Functor.ofOpSequence (functorMap f)).map (homOfLE h).op ≫ Pi.π _ k ≫ + eqToHom (functorObj_eq_neg (by omega : ¬(k < n))) = + (Pi.π (fun i ↦ if _ : i < m then M i else N i) k) ≫ + eqToHom (functorObj_eq_neg hh) := by + induction' h using Nat.leRec with m h ih + · simp + · specialize ih (by omega) + have : homOfLE (by omega : n ≤ m + 1) = + homOfLE (by omega : n ≤ m) ≫ homOfLE (by omega : m ≤ m + 1) := by simp + rw [this, op_comp, Functor.map_comp] + slice_lhs 2 4 => rw [ih] + simp only [Functor.ofOpSequence_obj, homOfLE_leOfHom, Functor.ofOpSequence_map_homOfLE_succ, + functorMap, dite_eq_ite, limMap_π_assoc, Discrete.functor_obj_eq_as, Discrete.natTrans_app] + split_ifs + simp [dif_neg (by omega : ¬(k < m))] + +lemma functorMap_commSq {n m : ℕ} (h : ¬(m < n)) : + (Functor.ofOpSequence (functorMap f)).map (homOfLE (by omega : n ≤ m + 1)).op ≫ Pi.π _ m ≫ + eqToHom (functorObj_eq_neg (by omega : ¬(m < n))) = + (Pi.π (fun i ↦ if _ : i < m + 1 then M i else N i) m) ≫ + eqToHom (functorObj_eq_pos (by omega)) ≫ f m := by + cases m with + | zero => + have : n = 0 := by omega + subst this + simp [functorMap] + | succ m => + rw [← functorMap_commSq_succ f (m + 1)] + simp only [Functor.ofOpSequence_obj, homOfLE_leOfHom, dite_eq_ite, + Functor.ofOpSequence_map_homOfLE_succ, add_le_iff_nonpos_right, nonpos_iff_eq_zero, + one_ne_zero] + have : homOfLE (by omega : n ≤ m + 1 + 1) = + homOfLE (by omega : n ≤ m + 1) ≫ homOfLE (by omega : m + 1 ≤ m + 1 + 1) := by simp + rw [this, op_comp, Functor.map_comp] + simp only [Functor.ofOpSequence_obj, homOfLE_leOfHom, Functor.ofOpSequence_map_homOfLE_succ, + Category.assoc, add_le_iff_nonpos_right, nonpos_iff_eq_zero, one_ne_zero] + congr 1 + exact functorMap_commSq_aux f (by omega) (by omega) + +/-- +The cone over the tower +``` +⋯ → ∏_{n < m} M n × ∏_{n ≥ m} N n → ⋯ → ∏ N +``` +with cone point `∏ M`. This is a limit cone, see `CategoryTheory.Limits.SequentialProduct.isLimit`. + -/ +noncomputable def cone : Cone (Functor.ofOpSequence (functorMap f)) where + pt := ∏ᶜ M + π := by + refine NatTrans.ofOpSequence + (fun n ↦ Limits.Pi.map fun m ↦ if h : m < n then eqToHom (functorObj_eq_pos h).symm else + f m ≫ eqToHom (functorObj_eq_neg h).symm) (fun n ↦ ?_) + apply Limits.Pi.hom_ext + intro m + simp only [Functor.const_obj_obj, Functor.ofOpSequence_obj, homOfLE_leOfHom, + Functor.const_obj_map, Category.id_comp, limMap_π, Discrete.functor_obj_eq_as, + Discrete.natTrans_app, Functor.ofOpSequence_map_homOfLE_succ, functorMap, Category.assoc, + limMap_π_assoc] + split + · simp [dif_pos (by omega : m < n + 1)] + · split + all_goals simp + +lemma cone_π_app (n : ℕ) : (cone f).π.app ⟨n⟩ = + Limits.Pi.map fun m ↦ if h : m < n then eqToHom (functorObj_eq_pos h).symm else + f m ≫ eqToHom (functorObj_eq_neg h).symm := rfl + +@[reassoc] +lemma cone_π_app_comp_Pi_π_pos (m n : ℕ) (h : n < m) : (cone f).π.app ⟨m⟩ ≫ + Pi.π (fun i ↦ if _ : i < m then M i else N i) n = + Pi.π _ n ≫ eqToHom (functorObj_eq_pos h).symm := by + simp only [Functor.const_obj_obj, dite_eq_ite, Functor.ofOpSequence_obj, cone_π_app, limMap_π, + Discrete.functor_obj_eq_as, Discrete.natTrans_app] + rw [dif_pos h] + +@[reassoc] +lemma cone_π_app_comp_Pi_π_neg (m n : ℕ) (h : ¬(n < m)) : (cone f).π.app ⟨m⟩ ≫ Pi.π _ n = + Pi.π _ n ≫ f n ≫ eqToHom (functorObj_eq_neg h).symm := by + simp only [Functor.const_obj_obj, dite_eq_ite, Functor.ofOpSequence_obj, cone_π_app, limMap_π, + Discrete.functor_obj_eq_as, Discrete.natTrans_app] + rw [dif_neg h] + +/-- +The cone over the tower +``` +⋯ → ∏_{n < m} M n × ∏_{n ≥ m} N n → ⋯ → ∏ N +``` +with cone point `∏ M` is indeed a limit cone. + -/ +noncomputable def isLimit : IsLimit (cone f) where + lift s := Pi.lift fun m ↦ + s.π.app ⟨m + 1⟩ ≫ Pi.π (fun i ↦ if _ : i < m + 1 then M i else N i) m ≫ + eqToHom (dif_pos (by omega : m < m + 1)) + fac s := by + intro ⟨n⟩ + apply Pi.hom_ext + intro m + by_cases h : m < n + · simp only [le_refl, Category.assoc, cone_π_app_comp_Pi_π_pos f _ _ h] + simp only [dite_eq_ite, Functor.ofOpSequence_obj, le_refl, limit.lift_π_assoc, Fan.mk_pt, + Discrete.functor_obj_eq_as, Fan.mk_π_app, Category.assoc, eqToHom_trans] + have hh : m + 1 ≤ n := by omega + rw [← s.w (homOfLE hh).op] + simp only [Functor.const_obj_obj, Functor.ofOpSequence_obj, homOfLE_leOfHom, le_refl, + Category.assoc] + congr + induction' hh using Nat.leRec with n hh ih + · simp + · have : homOfLE (Nat.le_succ_of_le hh) = homOfLE hh ≫ homOfLE (Nat.le_succ n) := by simp + rw [this, op_comp, Functor.map_comp] + simp only [Functor.ofOpSequence_obj, Nat.succ_eq_add_one, homOfLE_leOfHom, + Functor.ofOpSequence_map_homOfLE_succ, le_refl, Category.assoc] + have h₁ : (if _ : m < m + 1 then M m else N m) = if _ : m < n then M m else N m := by + rw [dif_pos (by omega), dif_pos (by omega)] + have h₂ : (if _ : m < n then M m else N m) = if _ : m < n + 1 then M m else N m := by + rw [dif_pos h, dif_pos (by omega)] + rw [← eqToHom_trans h₁ h₂] + slice_lhs 2 4 => rw [ih (by omega)] + simp only [functorMap, dite_eq_ite, Pi.π, limMap_π_assoc, Discrete.functor_obj_eq_as, + Discrete.natTrans_app] + split_ifs + rw [dif_pos (by omega)] + simp + · simp only [le_refl, Category.assoc] + rw [cone_π_app_comp_Pi_π_neg f _ _ h] + simp only [dite_eq_ite, Functor.ofOpSequence_obj, le_refl, limit.lift_π_assoc, Fan.mk_pt, + Discrete.functor_obj_eq_as, Fan.mk_π_app, Category.assoc] + slice_lhs 2 4 => erw [← functorMap_commSq f h] + simp + uniq s m h := by + apply Pi.hom_ext + intro n + simp only [Functor.ofOpSequence_obj, le_refl, dite_eq_ite, limit.lift_π, Fan.mk_pt, + Fan.mk_π_app, ← h ⟨n + 1⟩, Category.assoc] + slice_rhs 2 3 => erw [cone_π_app_comp_Pi_π_pos f (n + 1) _ (by omega)] + simp + +section + +variable [HasZeroMorphisms C] [HasFiniteBiproducts C] [HasCountableProducts C] [∀ n, Epi (f n)] + +attribute [local instance] hasBinaryBiproducts_of_finite_biproducts + +lemma functorMap_epi (n : ℕ) : Epi (functorMap f n) := by + rw [functorMap, Pi.map_eq_prod_map (P := fun m : ℕ ↦ m < n + 1)] + apply ( config := { allowSynthFailures := true } ) epi_comp + apply ( config := { allowSynthFailures := true } ) epi_comp + apply ( config := { allowSynthFailures := true } ) prod.map_epi + · apply ( config := { allowSynthFailures := true } ) Pi.map_epi + intro ⟨_, _⟩ + split + all_goals infer_instance + · apply ( config := { allowSynthFailures := true } ) IsIso.epi_of_iso + apply ( config := { allowSynthFailures := true } ) Pi.map_isIso + intro ⟨_, _⟩ + split + all_goals infer_instance +end + +end CategoryTheory.Limits.SequentialProduct diff --git a/Mathlib/Condensed/Light/AB.lean b/Mathlib/Condensed/Light/AB.lean new file mode 100644 index 0000000000000..ce4794aa35c4e --- /dev/null +++ b/Mathlib/Condensed/Light/AB.lean @@ -0,0 +1,30 @@ +/- +Copyright (c) 2024 Dagur Asgeirsson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Dagur Asgeirsson +-/ +import Mathlib.CategoryTheory.Abelian.GrothendieckAxioms +import Mathlib.Condensed.Light.Epi +import Mathlib.Condensed.Light.Limits +/-! + +# Grothendieck's AB axioms for light condensed modules + +The category of light condensed `R`-modules over a ring satisfies the countable version of +Grothendieck's AB4* axiom +-/ +universe u + +open CategoryTheory Limits + +namespace LightCondensed + +variable {R : Type u} [Ring R] + +attribute [local instance] Abelian.hasFiniteBiproducts + +noncomputable instance : CountableAB4Star (LightCondMod.{u} R) := + have := hasExactLimitsOfShape_of_preservesEpi (LightCondMod R) (Discrete ℕ) + CountableAB4Star.of_hasExactLimitsOfShape_nat _ + +end LightCondensed diff --git a/Mathlib/Condensed/Light/Epi.lean b/Mathlib/Condensed/Light/Epi.lean index ea2e380a1fbb1..795f7e9165f4b 100644 --- a/Mathlib/Condensed/Light/Epi.lean +++ b/Mathlib/Condensed/Light/Epi.lean @@ -3,10 +3,9 @@ Copyright (c) 2024 Dagur Asgeirsson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Dagur Asgeirsson -/ -import Mathlib.CategoryTheory.ConcreteCategory.EpiMono -import Mathlib.CategoryTheory.Sites.EpiMono +import Mathlib.CategoryTheory.Limits.Shapes.SequentialProduct import Mathlib.CategoryTheory.Sites.Coherent.SequentialLimit -import Mathlib.Condensed.Light.Module +import Mathlib.Condensed.Light.Limits /-! # Epimorphisms of light condensed objects @@ -14,6 +13,8 @@ import Mathlib.Condensed.Light.Module This file characterises epimorphisms in light condensed sets and modules as the locally surjective morphisms. Here, the condition of locally surjective is phrased in terms of continuous surjections of light profinite sets. + +Further, we prove that the functor `lim : Discrete ℕ ⥤ LightCondMod R` preserves epimorphisms. -/ universe v u w u' v' @@ -93,3 +94,35 @@ lemma epi_π_app_zero_of_epi : Epi (c.π.app ⟨0⟩) := by · exact fun _ ↦ (forget R).map_epi _ end LightCondensed + +open CategoryTheory.Limits.SequentialProduct + +namespace LightCondensed + +variable (n : ℕ) + +attribute [local instance] functorMap_epi Abelian.hasFiniteBiproducts + +variable {R : Type u} [Ring R] {M N : ℕ → LightCondMod.{u} R} (f : ∀ n, M n ⟶ N n) [∀ n, Epi (f n)] + +instance : Epi (Limits.Pi.map f) := by + have : Limits.Pi.map f = (cone f).π.app ⟨0⟩ := rfl + rw [this] + exact epi_π_app_zero_of_epi R (isLimit f) (fun n ↦ by simp; infer_instance) + +instance : (lim (J := Discrete ℕ) (C := LightCondMod R)).PreservesEpimorphisms where + preserves f _ := by + have : lim.map f = (Pi.isoLimit _).inv ≫ Limits.Pi.map (f.app ⟨·⟩) ≫ (Pi.isoLimit _).hom := by + apply limit.hom_ext + intro ⟨n⟩ + simp only [lim_obj, lim_map, limMap, IsLimit.map, limit.isLimit_lift, limit.lift_π, + Cones.postcompose_obj_pt, limit.cone_x, Cones.postcompose_obj_π, NatTrans.comp_app, + Functor.const_obj_obj, limit.cone_π, Pi.isoLimit, Limits.Pi.map, Category.assoc, + limit.conePointUniqueUpToIso_hom_comp, Pi.cone_pt, Pi.cone_π, Discrete.natTrans_app, + Discrete.functor_obj_eq_as] + erw [IsLimit.conePointUniqueUpToIso_inv_comp_assoc] + rfl + rw [this] + infer_instance + +end LightCondensed diff --git a/Mathlib/Condensed/Light/Limits.lean b/Mathlib/Condensed/Light/Limits.lean index c1509f55c090b..0955dd8d52178 100644 --- a/Mathlib/Condensed/Light/Limits.lean +++ b/Mathlib/Condensed/Light/Limits.lean @@ -23,8 +23,10 @@ instance : HasFiniteLimits LightCondSet.{u} := hasFiniteLimits_of_hasLimitsOfSiz variable (R : Type u) [Ring R] -instance : HasLimitsOfSize.{u, u} (LightCondMod.{u} R) := by - change HasLimitsOfSize (Sheaf _ _) - infer_instance +instance : HasLimitsOfSize.{u, u} (LightCondMod.{u} R) := + inferInstanceAs (HasLimitsOfSize (Sheaf _ _)) + +instance : HasLimitsOfSize.{0, 0} (LightCondMod.{u} R) := + inferInstanceAs (HasLimitsOfSize (Sheaf _ _)) instance : HasFiniteLimits (LightCondMod.{u} R) := hasFiniteLimits_of_hasLimitsOfSize _ From b244eb06a27d859da6147dce0a86d8fccd08f0b7 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 11 Dec 2024 21:35:51 +0000 Subject: [PATCH 671/829] feat: test that an executable with `import Mathlib` can be built and run (#19755) --- .../workflows/{lean4checker.yml => daily.yml} | 55 ++++++++++++++----- MathlibTest/MathlibTestExecutable.lean | 11 ++++ lakefile.lean | 3 + 3 files changed, 56 insertions(+), 13 deletions(-) rename .github/workflows/{lean4checker.yml => daily.yml} (67%) create mode 100644 MathlibTest/MathlibTestExecutable.lean diff --git a/.github/workflows/lean4checker.yml b/.github/workflows/daily.yml similarity index 67% rename from .github/workflows/lean4checker.yml rename to .github/workflows/daily.yml index c521c9164c994..f1f7dbe840941 100644 --- a/.github/workflows/lean4checker.yml +++ b/.github/workflows/daily.yml @@ -1,4 +1,7 @@ -name: lean4checker Workflow +name: Daily CI Workflow +# This workflow runs daily on `master` and the latest `nightly-testing-YYYY-MM-DD` tag, +# running some expensive CI checks that we don't want to run on every PR. +# It reports results via Zulip. on: schedule: @@ -41,14 +44,22 @@ jobs: LATEST_TAG=$(git tag | grep -E "${{ env.TAG_PATTERN }}" | sort -r | head -n 1) echo "LATEST_TAG=${LATEST_TAG}" >> "$GITHUB_ENV" + - name: Set branch ref + run: | + if [ "${{ matrix.branch_type }}" == "master" ]; then + echo "BRANCH_REF=${{ env.DEFAULT_BRANCH }}" >> "$GITHUB_ENV" + else + echo "BRANCH_REF=${{ env.LATEST_TAG }}" >> "$GITHUB_ENV" + fi + - name: Checkout branch or tag uses: actions/checkout@v4 with: - ref: ${{ matrix.branch_type == 'master' && env.DEFAULT_BRANCH || env.LATEST_TAG }} + ref: ${{ env.BRANCH_REF }} - name: If using a lean-pr-release toolchain, uninstall run: | - if [[ $(cat lean-toolchain) =~ ^leanprover/lean4-pr-releases:pr-release-[0-9]+$ ]]; then + if [[ "$(cat lean-toolchain)" =~ ^leanprover/lean4-pr-releases:pr-release-[0-9]+$ ]]; then printf 'Uninstalling transient toolchain %s\n' "$(cat lean-toolchain)" elan toolchain uninstall "$(cat lean-toolchain)" fi @@ -86,8 +97,13 @@ jobs: # so we explicitly check Batteries as well here. lake env lean4checker/.lake/build/bin/lean4checker Batteries Mathlib - - name: Post success message on Zulip - if: success() + - name: Run mathlib_test_executable + id: mathlib-test + run: | + lake exe mathlib_test_executable + + - name: Post success message for lean4checker on Zulip + if: steps.lean4checker.outcome == 'success' uses: zulip/github-actions-zulip/send-message@v1 with: api-key: ${{ secrets.ZULIP_API_KEY }} @@ -97,10 +113,23 @@ jobs: type: 'stream' topic: 'lean4checker' content: | - ✅ lean4checker [succeeded](https://github.com/${{ github.repository }}/actions/runs/${{ github.run_id }}) on ${{ github.sha }} (branch: ${{ matrix.branch_type == 'master' && 'master' || env.LATEST_TAG }}) + ✅ lean4checker [succeeded](https://github.com/${{ github.repository }}/actions/runs/${{ github.run_id }}) on ${{ github.sha }} (branch: ${{ env.BRANCH_REF }}) + + - name: Post success message for mathlib_test_executable on Zulip + if: steps.mathlib-test.outcome == 'success' + uses: zulip/github-actions-zulip/send-message@v1 + with: + api-key: ${{ secrets.ZULIP_API_KEY }} + email: 'github-mathlib4-bot@leanprover.zulipchat.com' + organization-url: 'https://leanprover.zulipchat.com' + to: 'mathlib reviewers' + type: 'stream' + topic: 'mathlib test executable' + content: | + ✅ mathlib_test_executable [succeeded](https://github.com/${{ github.repository }}/actions/runs/${{ github.run_id }}) on ${{ github.sha }} (branch: ${{ env.BRANCH_REF }}) - - name: Post failure message on Zulip - if: failure() + - name: Post failure message for lean4checker on Zulip + if: steps.lean4checker.outcome == 'failure' uses: zulip/github-actions-zulip/send-message@v1 with: api-key: ${{ secrets.ZULIP_API_KEY }} @@ -110,11 +139,11 @@ jobs: type: 'stream' topic: 'lean4checker failure' content: | - ❌ lean4checker [failed](https://github.com/${{ github.repository }}/actions/runs/${{ github.run_id }}) on ${{ github.sha }} (branch: ${{ matrix.branch_type == 'master' && 'master' || env.LATEST_TAG }}) + ❌ lean4checker [failed](https://github.com/${{ github.repository }}/actions/runs/${{ github.run_id }}) on ${{ github.sha }} (branch: ${{ env.BRANCH_REF }}) continue-on-error: true - - name: Post failure message on Zulip main topic - if: failure() + - name: Post failure message for mathlib_test_executable on Zulip + if: steps.mathlib-test.outcome == 'failure' uses: zulip/github-actions-zulip/send-message@v1 with: api-key: ${{ secrets.ZULIP_API_KEY }} @@ -122,7 +151,7 @@ jobs: organization-url: 'https://leanprover.zulipchat.com' to: 'mathlib reviewers' type: 'stream' - topic: 'lean4checker' + topic: 'mathlib test executable failure' content: | - ❌ lean4checker [failed](https://github.com/${{ github.repository }}/actions/runs/${{ github.run_id }}) on ${{ github.sha }} (branch: ${{ matrix.branch_type == 'master' && 'master' || env.LATEST_TAG }}) + ❌ mathlib_test_executable [failed](https://github.com/${{ github.repository }}/actions/runs/${{ github.run_id }}) on ${{ github.sha }} (branch: ${{ env.BRANCH_REF }}) continue-on-error: true diff --git a/MathlibTest/MathlibTestExecutable.lean b/MathlibTest/MathlibTestExecutable.lean new file mode 100644 index 0000000000000..60b105314cc63 --- /dev/null +++ b/MathlibTest/MathlibTestExecutable.lean @@ -0,0 +1,11 @@ +import Lean +import Std +import Qq +import Batteries +import Aesop +import ProofWidgets +import Mathlib +import Plausible + +def main : IO Unit := do + IO.println "Verified that an executable importing all of Mathlib and its upstream dependencies can be built and executed." diff --git a/lakefile.lean b/lakefile.lean index 7f2926ad0e7c5..1bf2f20f098ea 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -152,6 +152,9 @@ lean_exe unused where -- Executables which import `Lake` must set `-lLake`. weakLinkArgs := #["-lLake"] +lean_exe mathlib_test_executable where + root := `MathlibTest.MathlibTestExecutable + /-! ## Other configuration -/ From a608095eeb63e987177b864f008097d95c70abfe Mon Sep 17 00:00:00 2001 From: David Renshaw Date: Wed, 11 Dec 2024 22:38:45 +0000 Subject: [PATCH 672/829] chore(Topology/PartitionOfUnity): shorten proof using ParitionOfUnity.le_one (#19899) Replaces 11 lines of proof with a single invocation of [`ParitionOfUnity.le_one`](https://github.com/leanprover-community/mathlib4/blob/4411562e4fd2d02aa31e413bc8e948829aa3a62d/Mathlib/Topology/PartitionOfUnity.lean#L171). I observed this (via `set_option trace.profiler true`) to lower the elaboration time of this theorem from 0.20 seconds to 0.08 seconds. This simplification was found by [`tryAtEachStep`](https://github.com/dwrensha/tryAtEachStep). --- Mathlib/Topology/PartitionOfUnity.lean | 13 +------------ 1 file changed, 1 insertion(+), 12 deletions(-) diff --git a/Mathlib/Topology/PartitionOfUnity.lean b/Mathlib/Topology/PartitionOfUnity.lean index b78236ba907c2..fc38f80a2920d 100644 --- a/Mathlib/Topology/PartitionOfUnity.lean +++ b/Mathlib/Topology/PartitionOfUnity.lean @@ -665,15 +665,4 @@ theorem exists_continuous_sum_one_of_isOpen_isCompact [T2Space X] [LocallyCompac exact hj)] rfl intro i x - refine ⟨f.nonneg i x, ?_⟩ - by_cases h0 : f i x = 0 - · rw [h0] - exact zero_le_one - rw [← Finset.sum_singleton (f · x) i] - apply le_trans _ (f.sum_le_one' x) - rw [finsum_eq_sum (f.toFun · x) (by exact toFinite (support (f.toFun · x)))] - simp only [Finite.toFinset_setOf, ne_eq] - gcongr with z hz - · exact fun j _ _ => f.nonneg j x - simp only [Finset.singleton_subset_iff, Finset.mem_filter, Finset.mem_univ, true_and] - exact h0 + exact ⟨f.nonneg i x, PartitionOfUnity.le_one f i x⟩ From fbed4ff06dc59388a44b44517dc459679cb5be47 Mon Sep 17 00:00:00 2001 From: Johannes Choo Date: Thu, 12 Dec 2024 02:45:55 +0000 Subject: [PATCH 673/829] =?UTF-8?q?feat(Rat):=20`(q=E2=82=81=20+=20q?= =?UTF-8?q?=E2=82=82).den=20=E2=88=A3=20q=E2=82=81.den.lcm=20q=E2=82=82.de?= =?UTF-8?q?n`=20(#19842)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit add a useful theorem `Rat.add_den_dvd_lcm` that claims that the denominator of the addition of two rational numbers divides the LCM of the denominators of its terms. A stronger result than `Rat.add_den_dvd` that is already present, which still remains useful in many cases. --- Mathlib/Data/Rat/Lemmas.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/Mathlib/Data/Rat/Lemmas.lean b/Mathlib/Data/Rat/Lemmas.lean index dd45de30bdee3..726fce350d89c 100644 --- a/Mathlib/Data/Rat/Lemmas.lean +++ b/Mathlib/Data/Rat/Lemmas.lean @@ -63,6 +63,16 @@ theorem den_mk (n d : ℤ) : (n /. d).den = if d = 0 then 1 else d.natAbs / n.gc simp [divInt, mkRat, Rat.normalize, Nat.succPNat, Int.sign, Int.gcd, if_neg (Nat.cast_add_one_ne_zero _), this] +theorem add_den_dvd_lcm (q₁ q₂ : ℚ) : (q₁ + q₂).den ∣ q₁.den.lcm q₂.den := by + rw [add_def, normalize_eq, Nat.div_dvd_iff_dvd_mul (Nat.gcd_dvd_right _ _) + (Nat.gcd_ne_zero_right (by simp)), ← Nat.gcd_mul_lcm, + mul_dvd_mul_iff_right (Nat.lcm_ne_zero (by simp) (by simp)), Nat.dvd_gcd_iff] + refine ⟨?_, dvd_mul_right _ _⟩ + rw [← Int.natCast_dvd_natCast, Int.dvd_natAbs] + apply Int.dvd_add + <;> apply dvd_mul_of_dvd_right <;> rw [Int.natCast_dvd_natCast] + <;> [exact Nat.gcd_dvd_right _ _; exact Nat.gcd_dvd_left _ _] + theorem add_den_dvd (q₁ q₂ : ℚ) : (q₁ + q₂).den ∣ q₁.den * q₂.den := by rw [add_def, normalize_eq] apply Nat.div_dvd_of_dvd From 2ec7e0af4b6f7cc28f5a3ab26c7522d84e756b8a Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 12 Dec 2024 03:19:46 +0000 Subject: [PATCH 674/829] chore(MeasurableSpace/Prod): rename a lemma (#19901) Also use `Prod.map` in the statement and golf a proof using `aesop`. --- .../Constructions/BorelSpace/Real.lean | 2 +- .../MeasureTheory/MeasurableSpace/Prod.lean | 7 +++++-- .../Probability/Kernel/Composition/Basic.lean | 20 +++++++------------ .../Kernel/Disintegration/StandardBorel.lean | 6 +++--- 4 files changed, 16 insertions(+), 19 deletions(-) diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean index 5ed34323264cf..ee0c7e7404fb5 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean @@ -481,7 +481,7 @@ lemma measurable_of_real_prod {f : EReal × β → γ} (h_real : Measurable fun p : ℝ × β ↦ f (p.1, p.2)) (h_bot : Measurable fun x ↦ f (⊥, x)) (h_top : Measurable fun x ↦ f (⊤, x)) : Measurable f := .of_union₃_range_cover (measurableEmbedding_prod_mk_left _) (measurableEmbedding_prod_mk_left _) - (measurableEmbedding_coe.prod_mk .id) (by simp [-univ_subset_iff, subset_def, EReal.forall]) + (measurableEmbedding_coe.prodMap .id) (by simp [-univ_subset_iff, subset_def, EReal.forall]) h_bot h_top h_real lemma measurable_of_real_real {f : EReal × EReal → β} diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Prod.lean b/Mathlib/MeasureTheory/MeasurableSpace/Prod.lean index 9dec6c3e39e9b..ce0ec4a3af7a0 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Prod.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Prod.lean @@ -69,10 +69,10 @@ lemma isPiSystem_prod : IsPiSystem (image2 (· ×ˢ ·) { s : Set α | MeasurableSet s } { t : Set β | MeasurableSet t }) := isPiSystem_measurableSet.prod isPiSystem_measurableSet -lemma MeasurableEmbedding.prod_mk {α β γ δ : Type*} {mα : MeasurableSpace α} +lemma MeasurableEmbedding.prodMap {α β γ δ : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} {mγ : MeasurableSpace γ} {mδ : MeasurableSpace δ} {f : α → β} {g : γ → δ} (hg : MeasurableEmbedding g) (hf : MeasurableEmbedding f) : - MeasurableEmbedding fun x : γ × α => (g x.1, f x.2) := by + MeasurableEmbedding (Prod.map g f) := by have h_inj : Function.Injective fun x : γ × α => (g x.fst, f x.snd) := by intro x y hxy rw [← @Prod.mk.eta _ _ x, ← @Prod.mk.eta _ _ y] @@ -97,6 +97,9 @@ lemma MeasurableEmbedding.prod_mk {α β γ δ : Type*} {mα : MeasurableSpace simp_rw [Set.image_iUnion] exact MeasurableSet.iUnion hg +@[deprecated (since := "2024-12-11")] +alias MeasurableEmbedding.prod_mk := MeasurableEmbedding.prodMap + lemma MeasurableEmbedding.prod_mk_left {β γ : Type*} [MeasurableSingletonClass α] {mβ : MeasurableSpace β} {mγ : MeasurableSpace γ} (x : α) {f : γ → β} (hf : MeasurableEmbedding f) : diff --git a/Mathlib/Probability/Kernel/Composition/Basic.lean b/Mathlib/Probability/Kernel/Composition/Basic.lean index a992511575a5d..8fd4daf2faf96 100644 --- a/Mathlib/Probability/Kernel/Composition/Basic.lean +++ b/Mathlib/Probability/Kernel/Composition/Basic.lean @@ -571,21 +571,15 @@ lemma compProd_add_right (μ : Kernel α β) (κ η : Kernel (α × β) γ) lemma comapRight_compProd_id_prod {δ : Type*} {mδ : MeasurableSpace δ} (κ : Kernel α β) [IsSFiniteKernel κ] (η : Kernel (α × β) γ) [IsSFiniteKernel η] {f : δ → γ} (hf : MeasurableEmbedding f) : - comapRight (κ ⊗ₖ η) (MeasurableEmbedding.id.prod_mk hf) = κ ⊗ₖ (comapRight η hf) := by + comapRight (κ ⊗ₖ η) (MeasurableEmbedding.id.prodMap hf) = κ ⊗ₖ (comapRight η hf) := by ext a t ht rw [comapRight_apply' _ _ _ ht, compProd_apply, compProd_apply ht] - swap; · exact (MeasurableEmbedding.id.prod_mk hf).measurableSet_image.mpr ht - refine lintegral_congr (fun b ↦ ?_) - simp only [id_eq, Set.mem_image, Prod.mk.injEq, Prod.exists] - rw [comapRight_apply'] - swap; · exact measurable_prod_mk_left ht - congr with x - simp only [Set.mem_setOf_eq, Set.mem_image] - constructor - · rintro ⟨b', c, h, rfl, rfl⟩ - exact ⟨c, h, rfl⟩ - · rintro ⟨c, h, rfl⟩ - exact ⟨b, c, h, rfl, rfl⟩ + · refine lintegral_congr fun b ↦ ?_ + rw [comapRight_apply'] + · congr with x + aesop + · exact measurable_prod_mk_left ht + · exact (MeasurableEmbedding.id.prodMap hf).measurableSet_image.mpr ht end CompositionProduct diff --git a/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean b/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean index 9f1b75f64926a..e0368a6da934d 100644 --- a/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean @@ -242,13 +242,13 @@ lemma compProd_fst_borelMarkovFromReal_eq_comapRight_compProd = map κ (Prod.map (id : β → β) (embeddingReal Ω))) : fst κ ⊗ₖ borelMarkovFromReal Ω η = comapRight (fst (map κ (Prod.map (id : β → β) (embeddingReal Ω))) ⊗ₖ η) - (MeasurableEmbedding.id.prod_mk (measurableEmbedding_embeddingReal Ω)) := by + (MeasurableEmbedding.id.prodMap (measurableEmbedding_embeddingReal Ω)) := by let e := embeddingReal Ω let he := measurableEmbedding_embeddingReal Ω let κ' := map κ (Prod.map (id : β → β) e) have hη' : fst κ' ⊗ₖ η = κ' := hη have h_prod_embed : MeasurableEmbedding (Prod.map (id : β → β) e) := - MeasurableEmbedding.id.prod_mk he + MeasurableEmbedding.id.prodMap he change fst κ ⊗ₖ borelMarkovFromReal Ω η = comapRight (fst κ' ⊗ₖ η) h_prod_embed rw [comapRight_compProd_id_prod _ _ he] have h_fst : fst κ' = fst κ := by @@ -295,7 +295,7 @@ lemma compProd_fst_borelMarkovFromReal (κ : Kernel α (β × Ω)) [IsSFiniteKer let κ' := map κ (Prod.map (id : β → β) e) have hη' : fst κ' ⊗ₖ η = κ' := hη have h_prod_embed : MeasurableEmbedding (Prod.map (id : β → β) e) := - MeasurableEmbedding.id.prod_mk he + MeasurableEmbedding.id.prodMap he have : κ = comapRight κ' h_prod_embed := by ext c t : 2 unfold κ' From 67e672d5aba8e1e7cd4528bcc344c1af623ef4dd Mon Sep 17 00:00:00 2001 From: "Filippo A. E. Nuccio" Date: Thu, 12 Dec 2024 05:59:31 +0000 Subject: [PATCH 675/829] feat(Counterexamples/DiscreteTopologyNonDiscreteUniformity): add a counterexample about discrete uniformity and Cauchy filters (#17417) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We construct a non-discrete uniformity on `ℕ` that induces the discrete topology and admits a non-principal Cauchy filter (actually, the identity sequence is Cauchy). --- Counterexamples.lean | 1 + ...DiscreteTopologyNonDiscreteUniformity.lean | 315 ++++++++++++++++++ 2 files changed, 316 insertions(+) create mode 100644 Counterexamples/DiscreteTopologyNonDiscreteUniformity.lean diff --git a/Counterexamples.lean b/Counterexamples.lean index 85f441c3cd5d5..d6c0c3af5e0a5 100644 --- a/Counterexamples.lean +++ b/Counterexamples.lean @@ -3,6 +3,7 @@ import Counterexamples.CharPZeroNeCharZero import Counterexamples.CliffordAlgebraNotInjective import Counterexamples.Cyclotomic105 import Counterexamples.DirectSumIsInternal +import Counterexamples.DiscreteTopologyNonDiscreteUniformity import Counterexamples.GameMultiplication import Counterexamples.Girard import Counterexamples.HomogeneousPrimeNotPrime diff --git a/Counterexamples/DiscreteTopologyNonDiscreteUniformity.lean b/Counterexamples/DiscreteTopologyNonDiscreteUniformity.lean new file mode 100644 index 0000000000000..1442b1357a652 --- /dev/null +++ b/Counterexamples/DiscreteTopologyNonDiscreteUniformity.lean @@ -0,0 +1,315 @@ +/- +Copyright (c) 2024 Filippo A. E. Nuccio. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Filippo A. E. Nuccio +-/ + +import Mathlib.Analysis.SpecificLimits.Basic +import Mathlib.Order.Interval.Set.Basic +import Mathlib.Topology.MetricSpace.Pseudo.Defs +import Mathlib.Topology.MetricSpace.Cauchy +import Mathlib.Topology.UniformSpace.Cauchy + +/-! +# Discrete uniformities and discrete topology +Exactly as different metrics can induce equivalent topologies on a space, it is possible that +different uniform structures (a notion that generalises that of a metric structure) induce the same +topology on a space. In this file we are concerned in particular with the *discrete topology*, +formalised using the class `DiscreteTopology`, and the *discrete uniformity*, that is the bottom +element of the lattice of uniformities on a type (see `bot_uniformity`). + +The theorem `discreteTopology_of_discrete_uniformity` shows that the topology induced by the +discrete uniformity is the discrete one, but it is well-known that the converse might not hold in +general, along the lines of the above discussion. We explicitely produce a metric and a uniform +structure on a space (on `ℕ`, actually) that are not discrete, yet induce the discrete topology. + +To check that a certain uniformity is not discrete, recall that once a type `α` is endowed with a +uniformity, it is possible to speak about `Cauchy` filters on `a` and it is quite easy to see that +if the uniformity on `a` is the discrete one, a filter is Cauchy if and only if it coincides with +the principal filter `𝓟 {x}` (see `Filter.principal`) for some `x : α`. This is the +declaration `UniformSpace.DiscreteUnif.eq_const_of_cauchy` in Mathlib. + +A special case of this result is the intuitive observation that a sequence `a : ℕ → ℕ` can be a +Cauchy sequence if and only if it is eventually constant: when claiming this equivalence, one is +implicitely endowing `ℕ` with the metric inherited from `ℝ`, that induces the discrete uniformity +on `ℕ`. Hence, the intuition suggesting that a Cauchy sequence, whose +terms are "closer and closer to each other", valued in `ℕ` must be eventually constant for +*topological* reasons, namely the fact that `ℕ` is a discrete topological space, is *wrong* in the +sense that the reason is intrinsically "metric". In particular, if a non-constant sequence (like +the identity `id : ℕ → ℕ` is Cauchy, then the uniformity is certainly *not discrete*. + +## The counterexamples + +We produce two counterexamples: in the first section `Metric` we construct a metric and in the +second section `SetPointUniformity` we construct a uniformity, explicitely as a filter on `ℕ × ℕ`. +They basically coincide, and the difference of the examples lies in their flavours. + +### The metric + +We begin by defining a metric on `ℕ` (see `dist_def`) that +1. Induces the discrete topology, as proven in `TopIsDiscrete`; +2. Is not the discrete metric, in particular because the identity is a Cauchy sequence, as proven +in `idIsCauchy` + +The definition is simply `dist m n = |2 ^ (- n : ℤ) - 2 ^ (- m : ℤ)|`, and I am grateful to +Anatole Dedecker for his suggestion. + +### The set-point theoretic uniformity +A uniformity on `ℕ` is a filter on `ℕ × ℕ` satisfying some properties: we define a sequence of +subsets `fundamentalEntourage n : (Set ℕ × ℕ)` (indexed by `n : ℕ`) and we observe it satisfies the +condition needed to be a basis of a filter: moreover, the filter generated by this basis satisfies +the condition for being a uniformity, and this is the uniformity we put on `ℕ`. + +For each `n`, the set `fundamentalEntourage n : Set (ℕ × ℕ)` consists of the `n+1` points +`{(0,0),(1,1)...(n,n)}` on the diagonal; together with the half plane `{(x,y) | n ≤ x ∧ n ≤ y}` + +That this collection can be used as a filter basis is proven in the definition `counterBasis` and +that the filter `counterBasis.filterBasis` is a uniformity is proven in the definition +`counterCoreUniformity`. + +This induces the discrete topology, as proven in `TopIsDiscrete` and the `atTop` filter is Cauchy +(see `atTopIsCauchy`): that this specializes to the statement that the identity sequence +`id : ℕ → ℕ` is Cauchy is proven in `idIsCauchy`. + +## Implementation details +Since most of the statements evolve around membership of explicit natural numbers (framed by some +inequality) to explicit subsets, many proofs are easily closed by `aesop` or `omega` or `linarith`. + +### References +* [N. Bourbaki, *General Topology*, Chapter II][bourbaki1966] +-/ + +open Set Function Filter Metric + +/- We remove the "usual" instances of (discrete) topological space and of (discrete) uniform space +from `ℕ`-/ +attribute [-instance] instTopologicalSpaceNat instUniformSpaceNat + +section Metric + + +noncomputable local instance : PseudoMetricSpace ℕ where + dist := fun n m ↦ |2 ^ (- n : ℤ) - 2 ^ (- m : ℤ)| + dist_self := by simp only [zpow_neg, zpow_natCast, sub_self, abs_zero, implies_true] + dist_comm := fun _ _ ↦ abs_sub_comm .. + dist_triangle := fun _ _ _ ↦ abs_sub_le .. + +@[simp] +lemma dist_def {n m : ℕ} : dist n m = |2 ^ (-n : ℤ) - 2 ^ (-m : ℤ)| := rfl + +lemma Int.eq_of_pow_sub_le {d : ℕ} {m n : ℤ} (hd1 : 1 < d) + (h : |(d : ℝ) ^ (-m) - d ^ (-n)| < d ^ (-n - 1)) : m = n := by + have hd0 : 0 < d := one_pos.trans hd1 + replace h : |(1 : ℝ) - d ^ (n - m)| < (d : ℝ)⁻¹ := by + rw [← mul_lt_mul_iff_of_pos_left (a := (d : ℝ) ^ (-n)) (zpow_pos _ _), + ← abs_of_nonneg (a := (d : ℝ) ^ (-n)) (le_of_lt <| zpow_pos _ _), ← abs_mul, mul_sub, mul_one, + ← zpow_add₀ <| Nat.cast_ne_zero.mpr (ne_of_gt hd0), sub_eq_add_neg (b := m), + neg_add_cancel_left, ← abs_neg, neg_sub, + abs_of_nonneg (a := (d : ℝ) ^ (-n)) (le_of_lt <| zpow_pos _ _), ← zpow_neg_one, + ← zpow_add₀ <| Nat.cast_ne_zero.mpr (ne_of_gt hd0), ← sub_eq_add_neg] + exact h + all_goals exact Nat.cast_pos'.mpr hd0 + by_cases H : (m : ℤ) ≤ n + · obtain ⟨a, ha⟩ := Int.eq_ofNat_of_zero_le (sub_nonneg.mpr H) + rw [ha, ← mul_lt_mul_iff_of_pos_left (a := (d : ℝ)) <| Nat.cast_pos'.mpr hd0, + mul_inv_cancel₀ <| Nat.cast_ne_zero.mpr (ne_of_gt hd0), + ← abs_of_nonneg (a := (d : ℝ)) <| Nat.cast_nonneg' d, ← abs_mul, + show |(d : ℝ) * (1 - |(d : ℝ)| ^ (a : ℤ))| = |(d : ℤ) * (1 - |(d : ℤ)| ^ a)| by + norm_cast, ← Int.cast_one (R := ℝ), Int.cast_lt, Int.abs_lt_one_iff, Int.mul_eq_zero, + sub_eq_zero, eq_comm (a := 1), pow_eq_one_iff_cases] at h + simp only [Nat.cast_eq_zero, ne_of_gt hd0, Nat.abs_cast, Nat.cast_eq_one, ne_of_gt hd1, + Int.reduceNeg, reduceCtorEq, false_and, or_self, or_false, false_or] at h + rwa [h, Nat.cast_zero, sub_eq_zero, eq_comm] at ha + · have h1 : (d : ℝ) ^ (n - m) ≤ 1 - (d : ℝ)⁻¹ := by + calc (d : ℝ) ^ (n - m) ≤ (d : ℝ)⁻¹ := by + rw [← zpow_neg_one] + apply zpow_right_mono₀ <| Nat.one_le_cast.mpr hd0 + linarith + _ ≤ 1 - (d : ℝ)⁻¹ := by + rw [inv_eq_one_div, one_sub_div <| Nat.cast_ne_zero.mpr (ne_of_gt hd0), + div_le_div_iff_of_pos_right <| Nat.cast_pos'.mpr hd0, le_sub_iff_add_le] + norm_cast + linarith [sub_lt_of_abs_sub_lt_right (a := (1 : ℝ)) (b := d ^ (n - m)) (c := d⁻¹) h] + +lemma ball_eq_singleton {n : ℕ} : Metric.ball n ((2 : ℝ) ^ (-n - 1 : ℤ)) = {n} := by + ext m + constructor + · zify [zpow_natCast, mem_ball, dist_def, mem_singleton_iff] + apply Int.eq_of_pow_sub_le one_lt_two + · intro H + rw [H, mem_ball, dist_self] + apply zpow_pos two_pos + + +theorem TopIsDiscrete : DiscreteTopology ℕ := by + apply singletons_open_iff_discrete.mp + intro + simpa only [← ball_eq_singleton] using isOpen_ball + +lemma idIsCauchy : CauchySeq (id : ℕ → ℕ) := by + rw [Metric.cauchySeq_iff] + refine fun ε ↦ Metric.cauchySeq_iff.mp + (@cauchySeq_of_le_geometric_two ℝ _ 1 (fun n ↦ 2 ^(-n : ℤ)) fun n ↦ ?_) ε + simp only [zpow_natCast, Nat.cast_add, Nat.cast_one, neg_add_rev, Int.reduceNeg, one_div] + rw [Real.dist_eq, zpow_add' <| Or.intro_left _ two_ne_zero] + calc |2 ^ (- n : ℤ) - 2 ^ (-1 : ℤ) * 2 ^ (- n : ℤ)| = + |(1 - (2 : ℝ)⁻¹) * 2 ^ (- n : ℤ)| := by rw [← one_sub_mul, zpow_neg_one] + _ = |2⁻¹ * 2 ^ (-(n : ℤ))| := by congr; rw [inv_eq_one_div 2, sub_half 1] + _ = 2⁻¹ / 2 ^ n := by rw [zpow_neg, abs_mul, abs_inv, abs_inv, inv_eq_one_div, + Nat.abs_ofNat, one_div, zpow_natCast, abs_pow, ← div_eq_mul_inv, Nat.abs_ofNat] + rfl + +end Metric + +section SetPointUniformity + +/- As the `instance PseudoMetricSpace ℕ` declared in the previous section was local, `ℕ` has no +topology at this point. We are going to define a non-discrete uniform structure (just using the +filter-based definition), that will endow it with a topology that we will eventually show to be +discrete. -/ + +/-- The fundamental entourages (index by `n : ℕ`) used to construct a basis of the uniformity: for +each `n`, the set `fundamentalEntourage n : Set (ℕ × ℕ)` consists of the `n+1` points +`{(0,0),(1,1)...(n,n)}` on the diagonal; together with the half plane `{(x,y) | n ≤ x ∧ n ≤ y}`-/ +def fundamentalEntourage (n : ℕ) : Set (ℕ × ℕ) := + (⋃ i : Icc 0 n, {((i : ℕ), (i : ℕ))}) ∪ Set.Ici (n , n) + +@[simp] +lemma fundamentalEntourage_ext (t : ℕ) (T : Set (ℕ × ℕ)) : fundamentalEntourage t = T ↔ + T = (⋃ i : Icc 0 t, {((i : ℕ), (i : ℕ))}) ∪ Set.Ici (t , t) := by + simpa only [fundamentalEntourage] using eq_comm + +lemma mem_range_fundamentalEntourage (S : Set (ℕ × ℕ)) : + S ∈ (range fundamentalEntourage) ↔ ∃ n, fundamentalEntourage n = S := by + simp only [Set.mem_range, Eq.symm] + +lemma mem_fundamentalEntourage (n : ℕ) (P : ℕ × ℕ) : P ∈ fundamentalEntourage n ↔ + (n ≤ P.1 ∧ n ≤ P.2) ∨ (P.1 = P.2) := by + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · simp only [fundamentalEntourage, iUnion_singleton_eq_range, mem_union, mem_range, + Subtype.exists, mem_Icc, zero_le, true_and, exists_prop', nonempty_prop, mem_Ici] at h + rcases h with h | h + · apply Or.inr + rw [((h.choose_spec).2).symm] + · exact Or.inl h + · simp only [iUnion_singleton_eq_range, mem_union, mem_range, Subtype.exists, mem_Icc, zero_le, + true_and, exists_prop', nonempty_prop, mem_Ici, fundamentalEntourage] + rcases h with h | h + · exact Or.inr h + · by_cases h_le : n ≤ P.1 + · exact Or.inr ⟨h_le, h ▸ h_le⟩ + · refine Or.inl ⟨P.1, ⟨le_of_lt <| not_le.mp h_le, congrArg _ h⟩⟩ + +/-- The collection `fundamentalEntourage` satisfies the axioms to be a basis for a filter on + `ℕ × ℕ` and gives rise to a term in the relevant type.-/ +def counterBasis : FilterBasis (ℕ × ℕ) where + sets := range fundamentalEntourage + nonempty := range_nonempty _ + inter_sets := by + intro S T hS hT + obtain ⟨s, hs⟩ := hS + obtain ⟨t, ht⟩ := hT + simp only [mem_range, subset_inter_iff, exists_exists_eq_and, fundamentalEntourage] + use max t s + refine ⟨fun ⟨P1, P2⟩ hP ↦ ?_, fun ⟨P1, P2⟩ hP ↦ ?_⟩ <;> + cases' hP with h h <;> + simp only [iUnion_singleton_eq_range, mem_range, Prod.mk.injEq, Subtype.exists, mem_Icc, + zero_le, le_max_iff, true_and, exists_and_left, exists_prop', nonempty_prop, + exists_eq_left] at h + · simpa only [← hs, mem_fundamentalEntourage] using Or.inr h.2 + · simpa only [← hs, mem_fundamentalEntourage] using Or.inl + ⟨le_trans (by omega) h.1, le_trans (by omega) h.2⟩ + · simpa only [← ht, mem_fundamentalEntourage] using Or.inr h.2 + · simp only [mem_Ici, Prod.mk_le_mk] at h + simpa only [← ht, mem_fundamentalEntourage] using Or.inl ⟨le_trans + (by omega) h.1, le_trans (by omega) h.2⟩ + +@[simp] +lemma mem_counterBasis_iff (S : Set (ℕ × ℕ)) : + S ∈ counterBasis ↔ S ∈ range fundamentalEntourage := by + dsimp [counterBasis] + rfl + +/-- The "crude" uniform structure, without topology, simply as a the filter generated by `Basis` +and satisfying the axioms for being a uniformity. We later extract the topology `counterTopology` +generated by it and bundle `counterCoreUniformity` and `counterTopology` in a uniform strucutre +on `ℕ`, proving in passing that `counterTopology = ⊥`-/ +def counterCoreUniformity : UniformSpace.Core ℕ := by + apply UniformSpace.Core.mkOfBasis counterBasis <;> + intro S hS + · obtain ⟨n, hn⟩ := hS + simp only [fundamentalEntourage_ext, iUnion_singleton_eq_range] at hn + simp only [hn, mem_union, mem_range, Prod.mk.injEq, and_self, Subtype.exists, mem_Icc, zero_le, + true_and, exists_prop', nonempty_prop, exists_eq_right, mem_Ici, Prod.mk_le_mk] + omega + · refine ⟨S, hS, ?_⟩ + obtain ⟨n, hn⟩ := hS + simp only [fundamentalEntourage_ext, iUnion_singleton_eq_range] at hn + simp only [hn, preimage_union, union_subset_iff] + constructor + · apply subset_union_of_subset_left (subset_of_eq _) + aesop + · apply subset_union_of_subset_right (subset_of_eq _) + aesop + · refine ⟨S, hS, ?_⟩ + obtain ⟨n, hn⟩ := hS + simp only [fundamentalEntourage_ext, iUnion_singleton_eq_range] at hn + simp only [hn, mem_union, mem_range, Prod.mk.injEq, Subtype.exists, mem_Icc, zero_le, true_and, + exists_and_left, exists_prop', nonempty_prop, exists_eq_left, mem_Ici, Prod.mk_le_mk] + rintro ⟨P1, P2⟩ ⟨m, h1, h2⟩ + simp only [mem_union, mem_range, Prod.mk.injEq, Subtype.exists, mem_Icc, zero_le, true_and, + exists_and_left, exists_prop', nonempty_prop, exists_eq_left, mem_Ici, Prod.mk_le_mk] at h1 h2 + aesop + +/--The topology on `ℕ` induced by the "crude" uniformity-/ +instance counterTopology: TopologicalSpace ℕ := counterCoreUniformity.toTopologicalSpace + +/-- The uniform structure on `ℕ` bundling together the "crude" uniformity and the topology-/ +instance counterUniformity: UniformSpace ℕ := UniformSpace.ofCore counterCoreUniformity + +lemma HasBasis_counterUniformity : + (uniformity ℕ).HasBasis (fun _ ↦ True) fundamentalEntourage := by + show counterCoreUniformity.uniformity.HasBasis (fun _ ↦ True) fundamentalEntourage + simp only [Filter.hasBasis_iff, exists_and_left, true_and] + intro T + refine ⟨fun ⟨s, ⟨⟨r, hr⟩, hs⟩⟩ ↦ ⟨r, subset_of_eq_of_subset hr hs⟩ , fun ⟨n, hn⟩ ↦ ?_⟩ + exact (@FilterBasis.mem_filter_iff _ counterBasis T).mpr ⟨fundamentalEntourage n, by simp, hn⟩ + +/-- A proof that the topology on `ℕ` induced by the "crude" uniformity `counterCoreUniformity` +(or by `counterUniformity` tout-court, since they are `defeq`) is discrete.-/ +theorem TopIsDiscrete' : DiscreteTopology ℕ := by + rw [discreteTopology_iff_nhds] + intro n + rw [nhds_eq_comap_uniformity'] + apply Filter.ext + intro S + simp only [Filter.mem_comap, Filter.mem_pure] + have := @Filter.HasBasis.mem_uniformity_iff _ _ _ _ _ HasBasis_counterUniformity + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · simp_rw [this] at h + obtain ⟨T, ⟨⟨i, ⟨-, h1⟩⟩, h2⟩⟩ := h + apply h2 (h1 _ _ _) + rw [mem_fundamentalEntourage] + aesop + · refine ⟨fundamentalEntourage (n + 1), ?_, ?_⟩ + · show fundamentalEntourage (n + 1) ∈ counterCoreUniformity.uniformity + exact @Filter.HasBasis.mem_of_mem (ℕ × ℕ) ℕ counterCoreUniformity.uniformity (fun _ ↦ True) + fundamentalEntourage (n + 1) HasBasis_counterUniformity trivial + · simp only [preimage_subset_iff, mem_fundamentalEntourage, add_le_iff_nonpos_right, + nonpos_iff_eq_zero, one_ne_zero, and_false, false_or] + exact fun _ a ↦ mem_of_eq_of_mem a h + +/- With respect to the above uniformity, the `atTop` filter is Cauchy; in particular, it is not of +the form `𝓟 {x}` for any `x`, although the topology is discrete. This implies in passing that this +uniformity is not discrete-/ +lemma atTopIsCauchy : Cauchy (atTop : Filter ℕ) := by + rw [HasBasis_counterUniformity.cauchy_iff] + refine ⟨atTop_neBot, fun i _ ↦ ?_⟩ + simp_rw [mem_fundamentalEntourage, mem_atTop_sets, ge_iff_le] + exact ⟨Ici i, ⟨⟨i, fun _ hb ↦ hb⟩, fun _ hx _ hy ↦ Or.inl ⟨hx, hy⟩⟩⟩ + +/-- We find the same result about the identity map found in `idIsCauchy`, without using any metric +structure. -/ +lemma idIsCauchy' : CauchySeq (id : ℕ → _) := ⟨map_neBot, cauchy_iff_le.mp atTopIsCauchy⟩ + +end SetPointUniformity From b18b0c87025ebd469227afa91b78f861cac466b8 Mon Sep 17 00:00:00 2001 From: David Renshaw Date: Thu, 12 Dec 2024 08:11:29 +0000 Subject: [PATCH 676/829] chore: deprecate upstreamed theorem List.cons_subperm_of_mem (#19904) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Deprecates `List.cons_subperm_of_mem`, which was upstreamed to `Batteries` as [`List.cons_subperm_of_not_mem_of_mem`](https://github.com/leanprover-community/batteries/blob/74dffd1a83cdd2969a31c9892b0517e7c6f50668/Batteries/Data/List/Perm.lean#L94) in https://github.com/leanprover-community/batteries/pull/89. The upstream version is in fact *stronger*, as it does not have a `Nodup l₁` parameter. This duplication was found by [`tryAtEachStep`](https://github.com/dwrensha/tryAtEachStep). --- Mathlib/Data/List/Perm/Subperm.lean | 28 ++++------------------------ 1 file changed, 4 insertions(+), 24 deletions(-) diff --git a/Mathlib/Data/List/Perm/Subperm.lean b/Mathlib/Data/List/Perm/Subperm.lean index 3266017ace4f2..8623532c2a660 100644 --- a/Mathlib/Data/List/Perm/Subperm.lean +++ b/Mathlib/Data/List/Perm/Subperm.lean @@ -53,30 +53,10 @@ lemma subperm_cons_self : l <+~ a :: l := ⟨l, Perm.refl _, sublist_cons_self _ protected alias ⟨subperm.of_cons, subperm.cons⟩ := subperm_cons -theorem cons_subperm_of_mem {a : α} {l₁ l₂ : List α} (d₁ : Nodup l₁) (h₁ : a ∉ l₁) (h₂ : a ∈ l₂) - (s : l₁ <+~ l₂) : a :: l₁ <+~ l₂ := by - rcases s with ⟨l, p, s⟩ - induction s generalizing l₁ with - | slnil => cases h₂ - | @cons r₁ r₂ b s' ih => - simp? at h₂ says simp only [mem_cons] at h₂ - cases' h₂ with e m - · subst b - exact ⟨a :: r₁, p.cons a, s'.cons₂ _⟩ - · rcases ih d₁ h₁ m p with ⟨t, p', s'⟩ - exact ⟨t, p', s'.cons _⟩ - | @cons₂ r₁ r₂ b _ ih => - have bm : b ∈ l₁ := p.subset <| mem_cons_self _ _ - have am : a ∈ r₂ := by - simp only [find?, mem_cons] at h₂ - exact h₂.resolve_left fun e => h₁ <| e.symm ▸ bm - rcases append_of_mem bm with ⟨t₁, t₂, rfl⟩ - have st : t₁ ++ t₂ <+ t₁ ++ b :: t₂ := by simp - rcases ih (d₁.sublist st) (mt (fun x => st.subset x) h₁) am - (Perm.cons_inv <| p.trans perm_middle) with - ⟨t, p', s'⟩ - exact - ⟨b :: t, (p'.cons b).trans <| (swap _ _ _).trans (perm_middle.symm.cons a), s'.cons₂ _⟩ +@[deprecated List.cons_subperm_of_not_mem_of_mem (since := "2024-12-11"), nolint unusedArguments] +theorem cons_subperm_of_mem {a : α} {l₁ l₂ : List α} (_ : Nodup l₁) (h₁ : a ∉ l₁) (h₂ : a ∈ l₂) + (s : l₁ <+~ l₂) : a :: l₁ <+~ l₂ := + cons_subperm_of_not_mem_of_mem h₁ h₂ s protected theorem Nodup.subperm (d : Nodup l₁) (H : l₁ ⊆ l₂) : l₁ <+~ l₂ := subperm_of_subset d H From 07b2399c02d91a83fc030d165d734dca8d92a806 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Thu, 12 Dec 2024 09:27:51 +0000 Subject: [PATCH 677/829] refactor: make `LinearMap.IsAdjointPair` accept bare functions as arguments (#19679) Co-authored-by: Oliver Nash <7734364+ocfnash@users.noreply.github.com> --- Mathlib/Algebra/Lie/SkewAdjoint.lean | 6 +- .../BilinearForm/Properties.lean | 134 ------------------ .../LinearAlgebra/Matrix/BilinearForm.lean | 21 --- .../LinearAlgebra/RootSystem/OfBilinear.lean | 6 +- Mathlib/LinearAlgebra/SesquilinearForm.lean | 51 ++++--- 5 files changed, 35 insertions(+), 183 deletions(-) diff --git a/Mathlib/Algebra/Lie/SkewAdjoint.lean b/Mathlib/Algebra/Lie/SkewAdjoint.lean index 4ba8facc11fda..a39027840baa8 100644 --- a/Mathlib/Algebra/Lie/SkewAdjoint.lean +++ b/Mathlib/Algebra/Lie/SkewAdjoint.lean @@ -45,8 +45,8 @@ theorem LinearMap.BilinForm.isSkewAdjoint_bracket {f g : Module.End R M} (hf : f ∈ B.skewAdjointSubmodule) (hg : g ∈ B.skewAdjointSubmodule) : ⁅f, g⁆ ∈ B.skewAdjointSubmodule := by rw [mem_skewAdjointSubmodule] at * - have hfg : IsAdjointPair B B (f * g) (g * f) := by rw [← neg_mul_neg g f]; exact hf.mul hg - have hgf : IsAdjointPair B B (g * f) (f * g) := by rw [← neg_mul_neg f g]; exact hg.mul hf + have hfg : IsAdjointPair B B (f * g) (g * f) := by rw [← neg_mul_neg g f]; exact hg.comp hf + have hgf : IsAdjointPair B B (g * f) (f * g) := by rw [← neg_mul_neg f g]; exact hf.comp hg change IsAdjointPair B B (f * g - g * f) (-(f * g - g * f)); rw [neg_sub] exact hfg.sub hgf @@ -61,7 +61,7 @@ variable {N : Type w} [AddCommGroup N] [Module R N] (e : N ≃ₗ[R] M) /-- An equivalence of modules with bilinear forms gives equivalence of Lie algebras of skew-adjoint endomorphisms. -/ def skewAdjointLieSubalgebraEquiv : - skewAdjointLieSubalgebra (B.compl₁₂ (↑e : N →ₗ[R] M) ↑e) ≃ₗ⁅R⁆ skewAdjointLieSubalgebra B := by + skewAdjointLieSubalgebra (B.compl₁₂ (e : N →ₗ[R] M) e) ≃ₗ⁅R⁆ skewAdjointLieSubalgebra B := by apply LieEquiv.ofSubalgebras _ _ e.lieConj ext f simp only [LieSubalgebra.mem_coe, Submodule.mem_map_equiv, LieSubalgebra.mem_map_submodule, diff --git a/Mathlib/LinearAlgebra/BilinearForm/Properties.lean b/Mathlib/LinearAlgebra/BilinearForm/Properties.lean index 36af793f9e1a0..47aa85348d3ba 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Properties.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Properties.lean @@ -154,140 +154,6 @@ theorem isAlt_zero : (0 : BilinForm R M).IsAlt := fun _ => rfl theorem isAlt_neg {B : BilinForm R₁ M₁} : (-B).IsAlt ↔ B.IsAlt := ⟨fun h => neg_neg B ▸ h.neg, IsAlt.neg⟩ -/-! ### Linear adjoints -/ - - -section LinearAdjoints - -variable (B) (F : BilinForm R M) -variable {M' : Type*} [AddCommMonoid M'] [Module R M'] -variable (B' : BilinForm R M') (f f' : M →ₗ[R] M') (g g' : M' →ₗ[R] M) - -/-- Given a pair of modules equipped with bilinear forms, this is the condition for a pair of -maps between them to be mutually adjoint. -/ -def IsAdjointPair := - ∀ ⦃x y⦄, B' (f x) y = B x (g y) - -variable {B B' f f' g g'} - -theorem IsAdjointPair.eq (h : IsAdjointPair B B' f g) : ∀ {x y}, B' (f x) y = B x (g y) := - @h - -theorem isAdjointPair_iff_compLeft_eq_compRight (f g : Module.End R M) : - IsAdjointPair B F f g ↔ F.compLeft f = B.compRight g := by - constructor <;> intro h - · ext x - simp only [compLeft_apply, compRight_apply] - apply h - · intro x y - rw [← compLeft_apply, ← compRight_apply] - rw [h] - -theorem isAdjointPair_zero : IsAdjointPair B B' 0 0 := fun x y => by - simp only [BilinForm.zero_left, BilinForm.zero_right, LinearMap.zero_apply] - -theorem isAdjointPair_id : IsAdjointPair B B 1 1 := fun _ _ => rfl - -theorem IsAdjointPair.add (h : IsAdjointPair B B' f g) (h' : IsAdjointPair B B' f' g') : - IsAdjointPair B B' (f + f') (g + g') := fun x y => by - rw [LinearMap.add_apply, LinearMap.add_apply, add_left, add_right, h, h'] - -variable {M₁' : Type*} [AddCommGroup M₁'] [Module R₁ M₁'] -variable {B₁' : BilinForm R₁ M₁'} {f₁ f₁' : M₁ →ₗ[R₁] M₁'} {g₁ g₁' : M₁' →ₗ[R₁] M₁} - -theorem IsAdjointPair.sub (h : IsAdjointPair B₁ B₁' f₁ g₁) (h' : IsAdjointPair B₁ B₁' f₁' g₁') : - IsAdjointPair B₁ B₁' (f₁ - f₁') (g₁ - g₁') := fun x y => by - rw [LinearMap.sub_apply, LinearMap.sub_apply, sub_left, sub_right, h, h'] - -variable {B₂' : BilinForm R M'} {f₂ : M →ₗ[R] M'} {g₂ : M' →ₗ[R] M} - -theorem IsAdjointPair.smul (c : R) (h : IsAdjointPair B B₂' f₂ g₂) : - IsAdjointPair B B₂' (c • f₂) (c • g₂) := fun x y => by - rw [LinearMap.smul_apply, LinearMap.smul_apply, smul_left, smul_right, h] - -variable {M'' : Type*} [AddCommMonoid M''] [Module R M''] -variable (B'' : BilinForm R M'') - -theorem IsAdjointPair.comp {f' : M' →ₗ[R] M''} {g' : M'' →ₗ[R] M'} (h : IsAdjointPair B B' f g) - (h' : IsAdjointPair B' B'' f' g') : IsAdjointPair B B'' (f'.comp f) (g.comp g') := fun x y => by - rw [LinearMap.comp_apply, LinearMap.comp_apply, h', h] - -theorem IsAdjointPair.mul {f g f' g' : Module.End R M} (h : IsAdjointPair B B f g) - (h' : IsAdjointPair B B f' g') : IsAdjointPair B B (f * f') (g' * g) := fun x y => by - rw [LinearMap.mul_apply, LinearMap.mul_apply, h, h'] - -variable (B B' B₁ B₂) (F₂ : BilinForm R M) - -/-- The condition for an endomorphism to be "self-adjoint" with respect to a pair of bilinear forms -on the underlying module. In the case that these two forms are identical, this is the usual concept -of self adjointness. In the case that one of the forms is the negation of the other, this is the -usual concept of skew adjointness. -/ -def IsPairSelfAdjoint (f : Module.End R M) := - IsAdjointPair B F f f - -/-- The set of pair-self-adjoint endomorphisms are a submodule of the type of all endomorphisms. -/ -def isPairSelfAdjointSubmodule : Submodule R (Module.End R M) where - carrier := { f | IsPairSelfAdjoint B₂ F₂ f } - zero_mem' := isAdjointPair_zero - add_mem' hf hg := hf.add hg - smul_mem' c _ h := h.smul c - -@[simp] -theorem mem_isPairSelfAdjointSubmodule (f : Module.End R M) : - f ∈ isPairSelfAdjointSubmodule B₂ F₂ ↔ IsPairSelfAdjoint B₂ F₂ f := Iff.rfl - -theorem isPairSelfAdjoint_equiv (e : M' ≃ₗ[R] M) (f : Module.End R M) : - IsPairSelfAdjoint B₂ F₂ f ↔ - IsPairSelfAdjoint (B₂.comp ↑e ↑e) (F₂.comp ↑e ↑e) (e.symm.conj f) := by - have hₗ : (F₂.comp ↑e ↑e).compLeft (e.symm.conj f) = (F₂.compLeft f).comp ↑e ↑e := by - ext - simp [LinearEquiv.symm_conj_apply] - have hᵣ : (B₂.comp ↑e ↑e).compRight (e.symm.conj f) = (B₂.compRight f).comp ↑e ↑e := by - ext - simp [LinearEquiv.conj_apply] - have he : Function.Surjective (⇑(↑e : M' →ₗ[R] M) : M' → M) := e.surjective - show BilinForm.IsAdjointPair _ _ _ _ ↔ BilinForm.IsAdjointPair _ _ _ _ - rw [isAdjointPair_iff_compLeft_eq_compRight, isAdjointPair_iff_compLeft_eq_compRight, hᵣ, - hₗ, comp_inj _ _ he he] - -/-- An endomorphism of a module is self-adjoint with respect to a bilinear form if it serves as an -adjoint for itself. -/ -def IsSelfAdjoint (f : Module.End R M) := - IsAdjointPair B B f f - -/-- An endomorphism of a module is skew-adjoint with respect to a bilinear form if its negation -serves as an adjoint. -/ -def IsSkewAdjoint (f : Module.End R₁ M₁) := - IsAdjointPair B₁ B₁ f (-f) - -theorem isSkewAdjoint_iff_neg_self_adjoint (f : Module.End R₁ M₁) : - B₁.IsSkewAdjoint f ↔ IsAdjointPair (-B₁) B₁ f f := - show (∀ x y, B₁ (f x) y = B₁ x ((-f) y)) ↔ ∀ x y, B₁ (f x) y = (-B₁) x (f y) by - simp only [LinearMap.neg_apply, BilinForm.neg_apply, BilinForm.neg_right] - -/-- The set of self-adjoint endomorphisms of a module with bilinear form is a submodule. (In fact -it is a Jordan subalgebra.) -/ -def selfAdjointSubmodule := - isPairSelfAdjointSubmodule B B - -@[simp] -theorem mem_selfAdjointSubmodule (f : Module.End R M) : - f ∈ B.selfAdjointSubmodule ↔ B.IsSelfAdjoint f := - Iff.rfl - -/-- The set of skew-adjoint endomorphisms of a module with bilinear form is a submodule. (In fact -it is a Lie subalgebra.) -/ -def skewAdjointSubmodule := - isPairSelfAdjointSubmodule (-B₁) B₁ - -@[simp] -theorem mem_skewAdjointSubmodule (f : Module.End R₁ M₁) : - f ∈ B₁.skewAdjointSubmodule ↔ B₁.IsSkewAdjoint f := by - rw [isSkewAdjoint_iff_neg_self_adjoint] - exact Iff.rfl - -end LinearAdjoints - end BilinForm namespace BilinForm diff --git a/Mathlib/LinearAlgebra/Matrix/BilinearForm.lean b/Mathlib/LinearAlgebra/Matrix/BilinearForm.lean index 6ba3a7550fedf..7c27921ccd099 100644 --- a/Mathlib/LinearAlgebra/Matrix/BilinearForm.lean +++ b/Mathlib/LinearAlgebra/Matrix/BilinearForm.lean @@ -297,20 +297,6 @@ variable {n : Type*} [Fintype n] variable (b : Basis n R₂ M₂) variable (J J₃ A A' : Matrix n n R₂) -@[simp] -theorem isAdjointPair_toBilin' [DecidableEq n] : - BilinForm.IsAdjointPair (Matrix.toBilin' J) (Matrix.toBilin' J₃) (Matrix.toLin' A) - (Matrix.toLin' A') ↔ - Matrix.IsAdjointPair J J₃ A A' := - isAdjointPair_toLinearMap₂' _ _ _ _ - -@[simp] -theorem isAdjointPair_toBilin [DecidableEq n] : - BilinForm.IsAdjointPair (Matrix.toBilin b J) (Matrix.toBilin b J₃) (Matrix.toLin b b A) - (Matrix.toLin b b A') ↔ - Matrix.IsAdjointPair J J₃ A A' := - isAdjointPair_toLinearMap₂ _ _ _ _ _ _ - theorem Matrix.isAdjointPair_equiv' [DecidableEq n] (P : Matrix n n R₂) (h : IsUnit P) : (Pᵀ * J * P).IsAdjointPair (Pᵀ * J * P) A A' ↔ J.IsAdjointPair J (P * A * P⁻¹) (P * A' * P⁻¹) := @@ -318,13 +304,6 @@ theorem Matrix.isAdjointPair_equiv' [DecidableEq n] (P : Matrix n n R₂) (h : I variable [DecidableEq n] -/-- The submodule of pair-self-adjoint matrices with respect to bilinear forms corresponding to -given matrices `J`, `J₂`. -/ -def pairSelfAdjointMatricesSubmodule' : Submodule R₂ (Matrix n n R₂) := - (BilinForm.isPairSelfAdjointSubmodule (Matrix.toBilin' J) (Matrix.toBilin' J₃)).map - ((LinearMap.toMatrix' : ((n → R₂) →ₗ[R₂] n → R₂) ≃ₗ[R₂] Matrix n n R₂) : - ((n → R₂) →ₗ[R₂] n → R₂) →ₗ[R₂] Matrix n n R₂) - theorem mem_pairSelfAdjointMatricesSubmodule' : A ∈ pairSelfAdjointMatricesSubmodule J J₃ ↔ Matrix.IsAdjointPair J J₃ A A := by simp only [mem_pairSelfAdjointMatricesSubmodule] diff --git a/Mathlib/LinearAlgebra/RootSystem/OfBilinear.lean b/Mathlib/LinearAlgebra/RootSystem/OfBilinear.lean index ae27ada2167b0..3b9f88b113467 100644 --- a/Mathlib/LinearAlgebra/RootSystem/OfBilinear.lean +++ b/Mathlib/LinearAlgebra/RootSystem/OfBilinear.lean @@ -90,13 +90,13 @@ lemma reflective_reflection (hSB : LinearMap.IsSymm B) {y : M} (hx : IsReflective B x) (hy : IsReflective B y) : IsReflective B (Module.reflection (coroot_apply_self B hx) y) := by constructor - · rw [← LinearEquiv.coe_coe, isOrthogonal_reflection B hx hSB] + · rw [isOrthogonal_reflection B hx hSB] exact hy.1 · intro z have hz : Module.reflection (coroot_apply_self B hx) (Module.reflection (coroot_apply_self B hx) z) = z := by exact (LinearEquiv.eq_symm_apply (Module.reflection (coroot_apply_self B hx))).mp rfl - rw [← hz, ← LinearEquiv.coe_coe, isOrthogonal_reflection B hx hSB, + rw [← hz, isOrthogonal_reflection B hx hSB, isOrthogonal_reflection B hx hSB] exact hy.2 _ @@ -170,7 +170,7 @@ def ofBilinear [IsReflexive R M] (B : M →ₗ[R] M →ₗ[R] R) (hNB : LinearMa simp only [mem_setOf_eq, PerfectPairing.flip_apply_apply, mul_sub, apply_self_mul_coroot_apply B y.2, ← mul_assoc] rw [← isOrthogonal_reflection B x.2 hSB y y, apply_self_mul_coroot_apply, ← hSB z, ← hSB z, - RingHom.id_apply, RingHom.id_apply, LinearEquiv.coe_coe, Module.reflection_apply, map_sub, + RingHom.id_apply, RingHom.id_apply, Module.reflection_apply, map_sub, mul_sub, sub_eq_sub_iff_comm, sub_left_inj] refine x.2.1.1 ?_ simp only [mem_setOf_eq, map_smul, smul_eq_mul] diff --git a/Mathlib/LinearAlgebra/SesquilinearForm.lean b/Mathlib/LinearAlgebra/SesquilinearForm.lean index 495286b640a80..19fb599ac6529 100644 --- a/Mathlib/LinearAlgebra/SesquilinearForm.lean +++ b/Mathlib/LinearAlgebra/SesquilinearForm.lean @@ -15,8 +15,8 @@ form `M₁ →ₛₗ[I₁] M₂ →ₛₗ[I₂] M`, where `I₁ : R₁ →+* R` Sesquilinear forms are the special case that `M₁ = M₂`, `M = R₁ = R₂ = R`, and `I₁ = RingHom.id R`. Taking additionally `I₂ = RingHom.id R`, then one obtains bilinear forms. -These forms are a special case of the bilinear maps defined in `BilinearMap.lean` and all basic -lemmas about construction and elementary calculations are found there. +Sesquilinear maps are a special case of the bilinear maps defined in `BilinearMap.lean` and `many` +basic lemmas about construction and elementary calculations are found there. ## Main declarations @@ -410,7 +410,7 @@ variable (B B' f g) /-- Given a pair of modules equipped with bilinear maps, this is the condition for a pair of maps between them to be mutually adjoint. -/ -def IsAdjointPair := +def IsAdjointPair (f : M → M₁) (g : M₁ → M) := ∀ x y, B' (f x) y = B x (g y) variable {B B' f g} @@ -423,17 +423,24 @@ theorem isAdjointPair_iff_comp_eq_compl₂ : IsAdjointPair B B' f g ↔ B'.comp · intro _ _ rw [← compl₂_apply, ← comp_apply, h] -theorem isAdjointPair_zero : IsAdjointPair B B' 0 0 := fun _ _ ↦ by simp only [zero_apply, map_zero] +theorem isAdjointPair_zero : IsAdjointPair B B' 0 0 := fun _ _ ↦ by + simp only [Pi.zero_apply, map_zero, zero_apply] -theorem isAdjointPair_id : IsAdjointPair B B 1 1 := fun _ _ ↦ rfl +theorem isAdjointPair_id : IsAdjointPair B B (_root_.id : M → M) (_root_.id : M → M) := + fun _ _ ↦ rfl -theorem IsAdjointPair.add (h : IsAdjointPair B B' f g) (h' : IsAdjointPair B B' f' g') : +theorem isAdjointPair_one : IsAdjointPair B B (1 : Module.End R M) (1 : Module.End R M) := + isAdjointPair_id + +theorem IsAdjointPair.add {f f' : M → M₁} {g g' : M₁ → M} (h : IsAdjointPair B B' f g) + (h' : IsAdjointPair B B' f' g') : IsAdjointPair B B' (f + f') (g + g') := fun x _ ↦ by - rw [f.add_apply, g.add_apply, B'.map_add₂, (B x).map_add, h, h'] + rw [Pi.add_apply, Pi.add_apply, B'.map_add₂, (B x).map_add, h, h'] -theorem IsAdjointPair.comp {f' : M₁ →ₗ[R] M₂} {g' : M₂ →ₗ[R] M₁} (h : IsAdjointPair B B' f g) - (h' : IsAdjointPair B' B'' f' g') : IsAdjointPair B B'' (f'.comp f) (g.comp g') := fun _ _ ↦ by - rw [LinearMap.comp_apply, LinearMap.comp_apply, h', h] +theorem IsAdjointPair.comp {f : M → M₁} {g : M₁ → M} {f' : M₁ → M₂} {g' : M₂ → M₁} + (h : IsAdjointPair B B' f g) (h' : IsAdjointPair B' B'' f' g') : + IsAdjointPair B B'' (f' ∘ f) (g ∘ g') := fun _ _ ↦ by + rw [Function.comp_def, Function.comp_def, h', h] theorem IsAdjointPair.mul {f g f' g' : Module.End R M} (h : IsAdjointPair B B f g) (h' : IsAdjointPair B B f' g') : IsAdjointPair B B (f * f') (g' * g) := @@ -448,11 +455,11 @@ variable [AddCommGroup M] [Module R M] variable [AddCommGroup M₁] [Module R M₁] variable [AddCommGroup M₂] [Module R M₂] variable {B F : M →ₗ[R] M →ₗ[R] M₂} {B' : M₁ →ₗ[R] M₁ →ₗ[R] M₂} -variable {f f' : M →ₗ[R] M₁} {g g' : M₁ →ₗ[R] M} +variable {f f' : M → M₁} {g g' : M₁ → M} theorem IsAdjointPair.sub (h : IsAdjointPair B B' f g) (h' : IsAdjointPair B B' f' g') : IsAdjointPair B B' (f - f') (g - g') := fun x _ ↦ by - rw [f.sub_apply, g.sub_apply, B'.map_sub₂, (B x).map_sub, h, h'] + rw [Pi.sub_apply, Pi.sub_apply, B'.map_sub₂, (B x).map_sub, h, h'] theorem IsAdjointPair.smul (c : R) (h : IsAdjointPair B B' f g) : IsAdjointPair B B' (c • f) (c • g) := fun _ _ ↦ by @@ -463,7 +470,7 @@ end AddCommGroup section OrthogonalMap variable {R M : Type*} [CommRing R] [AddCommGroup M] [Module R M] - (B : LinearMap.BilinForm R M) (f : Module.End R M) + (B : LinearMap.BilinForm R M) (f : M → M) /-- A linear transformation `f` is orthogonal with respect to a bilinear form `B` if `B` is bi-invariant with respect to `f`. -/ @@ -473,17 +480,17 @@ def IsOrthogonal : Prop := variable {B f} @[simp] -lemma _root_.LinearEquiv.isAdjointPair_symm_iff {f : M ≃ₗ[R] M} : +lemma _root_.LinearEquiv.isAdjointPair_symm_iff {f : M ≃ M} : LinearMap.IsAdjointPair B B f f.symm ↔ B.IsOrthogonal f := ⟨fun hf x y ↦ by simpa using hf x (f y), fun hf x y ↦ by simpa using hf x (f.symm y)⟩ -lemma isOrthogonal_of_forall_apply_same - (h : IsLeftRegular (2 : R)) (hB : B.IsSymm) (hf : ∀ x, B (f x) (f x) = B x x) : +lemma isOrthogonal_of_forall_apply_same {F : Type*} [FunLike F M M] [LinearMapClass F R M M] + (f : F) (h : IsLeftRegular (2 : R)) (hB : B.IsSymm) (hf : ∀ x, B (f x) (f x) = B x x) : B.IsOrthogonal f := by intro x y suffices 2 * B (f x) (f y) = 2 * B x y from h this have := hf (x + y) - simp only [map_add, add_apply, hf x, hf y, show B y x = B x y from hB.eq y x] at this + simp only [map_add, LinearMap.add_apply, hf x, hf y, show B y x = B x y from hB.eq y x] at this rw [show B (f y) (f x) = B (f x) (f y) from hB.eq (f y) (f x)] at this simp only [add_assoc, add_right_inj] at this simp only [← add_assoc, add_left_inj] at this @@ -510,12 +517,12 @@ variable (B F : M →ₗ[R] M →ₛₗ[I] M₁) on the underlying module. In the case that these two maps are identical, this is the usual concept of self adjointness. In the case that one of the maps is the negation of the other, this is the usual concept of skew adjointness. -/ -def IsPairSelfAdjoint (f : Module.End R M) := +def IsPairSelfAdjoint (f : M → M) := IsAdjointPair B F f f /-- An endomorphism of a module is self-adjoint with respect to a bilinear map if it serves as an adjoint for itself. -/ -protected def IsSelfAdjoint (f : Module.End R M) := +protected def IsSelfAdjoint (f : M → M) := IsAdjointPair B B f f end AddCommMonoid @@ -535,7 +542,7 @@ def isPairSelfAdjointSubmodule : Submodule R (Module.End R M) where /-- An endomorphism of a module is skew-adjoint with respect to a bilinear map if its negation serves as an adjoint. -/ -def IsSkewAdjoint (f : Module.End R M) := +def IsSkewAdjoint (f : M → M) := IsAdjointPair B B f (-f) /-- The set of self-adjoint endomorphisms of a module with bilinear map is a submodule. (In fact @@ -557,7 +564,7 @@ theorem mem_isPairSelfAdjointSubmodule (f : Module.End R M) : theorem isPairSelfAdjoint_equiv (e : M₁ ≃ₗ[R] M) (f : Module.End R M) : IsPairSelfAdjoint B F f ↔ - IsPairSelfAdjoint (B.compl₁₂ ↑e ↑e) (F.compl₁₂ ↑e ↑e) (e.symm.conj f) := by + IsPairSelfAdjoint (B.compl₁₂ e e) (F.compl₁₂ e e) (e.symm.conj f) := by have hₗ : (F.compl₁₂ (↑e : M₁ →ₗ[R] M) (↑e : M₁ →ₗ[R] M)).comp (e.symm.conj f) = (F.comp f).compl₁₂ (↑e : M₁ →ₗ[R] M) (↑e : M₁ →ₗ[R] M) := by @@ -573,7 +580,7 @@ theorem isPairSelfAdjoint_equiv (e : M₁ ≃ₗ[R] M) (f : Module.End R M) : have he : Function.Surjective (⇑(↑e : M₁ →ₗ[R] M) : M₁ → M) := e.surjective simp_rw [IsPairSelfAdjoint, isAdjointPair_iff_comp_eq_compl₂, hₗ, hᵣ, compl₁₂_inj he he] -theorem isSkewAdjoint_iff_neg_self_adjoint (f : Module.End R M) : +theorem isSkewAdjoint_iff_neg_self_adjoint (f : M → M) : B.IsSkewAdjoint f ↔ IsAdjointPair (-B) B f f := show (∀ x y, B (f x) y = B x ((-f) y)) ↔ ∀ x y, B (f x) y = (-B) x (f y) by simp From 415efa61fc0b3ee799f90fe287954d6016bd1bac Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Thu, 12 Dec 2024 11:32:08 +0000 Subject: [PATCH 678/829] feat(AlgebraicTopology): define the simplicial nerve of a simplicial category (#19837) This PR defines the simplicial (/homotopy coherent) nerve of a simplicial category. --- Mathlib.lean | 1 + .../AlgebraicTopology/SimplicialNerve.lean | 228 ++++++++++++++++++ .../SimplicialSet/Basic.lean | 4 + .../SimplicialSet/Nerve.lean | 14 +- Mathlib/CategoryTheory/Enriched/Basic.lean | 10 + Mathlib/Order/Hom/Basic.lean | 6 + docs/references.bib | 7 + 7 files changed, 269 insertions(+), 1 deletion(-) create mode 100644 Mathlib/AlgebraicTopology/SimplicialNerve.lean diff --git a/Mathlib.lean b/Mathlib.lean index 236d70218b0b1..8d15eac376c50 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1053,6 +1053,7 @@ import Mathlib.AlgebraicTopology.Quasicategory.StrictSegal import Mathlib.AlgebraicTopology.SimplexCategory import Mathlib.AlgebraicTopology.SimplicialCategory.Basic import Mathlib.AlgebraicTopology.SimplicialCategory.SimplicialObject +import Mathlib.AlgebraicTopology.SimplicialNerve import Mathlib.AlgebraicTopology.SimplicialObject.Basic import Mathlib.AlgebraicTopology.SimplicialObject.Coskeletal import Mathlib.AlgebraicTopology.SimplicialSet.Basic diff --git a/Mathlib/AlgebraicTopology/SimplicialNerve.lean b/Mathlib/AlgebraicTopology/SimplicialNerve.lean new file mode 100644 index 0000000000000..9cb1dece1a5cf --- /dev/null +++ b/Mathlib/AlgebraicTopology/SimplicialNerve.lean @@ -0,0 +1,228 @@ +/- +Copyright (c) 2024 Dagur Asgeirsson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Dagur Asgeirsson +-/ +import Mathlib.AlgebraicTopology.SimplicialCategory.Basic +import Mathlib.AlgebraicTopology.SimplicialSet.Nerve +/-! + +# The simplicial nerve of a simplicial category + +This file defines the simplicial nerve (sometimes called homotopy coherent nerve) of a simplicial +category. + +We define the *simplicial thickening* of a linear order `J` as the simplicial category whose hom +objects `i ⟶ j` are given by the nerve of the poset of "paths" from `i` to `j` in `J`. This is the +poset of subsets of the interval `[i, j]` in `J`, containing the endpoints. + +The simplicial nerve of a simplicial category `C` is then defined as the simplicial set whose +`n`-simplices are given by the set of simplicial functors from the simplicial thickening of +the linear order `Fin (n + 1)` to `C`, in other words +`SimplicialNerve C _[n] := EnrichedFunctor SSet (SimplicialThickening (Fin (n + 1))) C`. + +## Projects + +* Prove that the 0-simplicies of `SimplicialNerve C` may be identified with the objects of `C` +* Prove that the 1-simplicies of `SimplicialNerve C` may be identified with the morphisms of `C` +* Prove that the simplicial nerve of a simplicial category `C`, such that `sHom X Y` is a Kan + complex for every pair of objects `X Y : C`, is a quasicategory. +* Define the quasicategory of anima as the simplicial nerve of the simplicial category of + Kan complexes. +* Define the functor from topological spaces to anima. + +## References +* [Jacob Lurie, *Higher Topos Theory*, Section 1.1.5][LurieHTT] +-/ + +universe v u + +namespace CategoryTheory + +open SimplicialCategory EnrichedCategory EnrichedOrdinaryCategory MonoidalCategory + +open scoped Simplicial + +section SimplicialNerve + +/-- A type synonym for a linear order `J`, will be equipped with a simplicial category structure. -/ +@[nolint unusedArguments] +def SimplicialThickening (J : Type*) [LinearOrder J] : Type _ := J + +instance (J : Type*) [LinearOrder J] : LinearOrder (SimplicialThickening J) := + inferInstanceAs (LinearOrder J) + +namespace SimplicialThickening + +/-- +A path from `i` to `j` in a linear order `J` is a subset of the interval `[i, j]` in `J` containing +the endpoints. +-/ +@[ext] +structure Path {J : Type*} [LinearOrder J] (i j : J) where + /-- The underlying subset -/ + I : Set J + left : i ∈ I := by simp + right : j ∈ I := by simp + left_le (k : J) (_ : k ∈ I) : i ≤ k := by simp + le_right (k : J) (_ : k ∈ I) : k ≤ j := by simp + +lemma Path.le {J : Type*} [LinearOrder J] {i j : J} (f : Path i j) : i ≤ j := + f.left_le _ f.right + +instance {J : Type*} [LinearOrder J] (i j : J) : Category (Path i j) := + InducedCategory.category (fun f : Path i j ↦ f.I) + +@[simps] +instance (J : Type*) [LinearOrder J] : CategoryStruct (SimplicialThickening J) where + Hom i j := Path i j + id i := { I := {i} } + comp {i j k} f g := { + I := f.I ∪ g.I + left := Or.inl f.left + right := Or.inr g.right + left_le l := by + rintro (h | h) + exacts [(f.left_le l h), (Path.le f).trans (g.left_le l h)] + le_right l := by + rintro (h | h) + exacts [(f.le_right _ h).trans (Path.le g), (g.le_right l h)] } + +instance {J : Type*} [LinearOrder J] (i j : SimplicialThickening J) : Category (i ⟶ j) := + inferInstanceAs (Category (Path _ _)) + +@[ext] +lemma hom_ext {J : Type*} [LinearOrder J] + (i j : SimplicialThickening J) (x y : i ⟶ j) (h : ∀ t, t ∈ x.I ↔ t ∈ y.I) : x = y := by + apply Path.ext + ext + apply h + +instance (J : Type*) [LinearOrder J] : Category (SimplicialThickening J) where + id_comp f := by ext; simpa using fun h ↦ h ▸ f.left + comp_id f := by ext; simpa using fun h ↦ h ▸ f.right + +/-- +Composition of morphisms in `SimplicialThickening J`, as a functor `(i ⟶ j) × (j ⟶ k) ⥤ (i ⟶ k)` +-/ +@[simps] +def compFunctor {J : Type*} [LinearOrder J] + (i j k : SimplicialThickening J) : (i ⟶ j) × (j ⟶ k) ⥤ (i ⟶ k) where + obj x := x.1 ≫ x.2 + map f := ⟨⟨Set.union_subset_union f.1.1.1 f.2.1.1⟩⟩ + +namespace SimplicialCategory + +variable {J : Type*} [LinearOrder J] + +/-- The hom simplicial set of the simplicial category structure on `SimplicialThickening J` -/ +abbrev Hom (i j : SimplicialThickening J) : SSet := (nerve (i ⟶ j)) + +/-- The identity of the simplicial category structure on `SimplicialThickening J` -/ +abbrev id (i : SimplicialThickening J) : 𝟙_ SSet ⟶ Hom i i := + ⟨fun _ _ ↦ (Functor.const _).obj (𝟙 _), fun _ _ _ ↦ by simp; rfl⟩ + +/-- The composition of the simplicial category structure on `SimplicialThickening J` -/ +abbrev comp (i j k : SimplicialThickening J) : Hom i j ⊗ Hom j k ⟶ Hom i k := + ⟨fun _ x ↦ x.1.prod' x.2 ⋙ compFunctor i j k, fun _ _ _ ↦ by simp; rfl⟩ + +@[simp] +lemma id_comp (i j : SimplicialThickening J) : + (λ_ (Hom i j)).inv ≫ id i ▷ Hom i j ≫ comp i i j = 𝟙 (Hom i j) := by + rw [Iso.inv_comp_eq] + ext + exact Functor.ext (fun _ ↦ by simp) + +@[simp] +lemma comp_id (i j : SimplicialThickening J) : + (ρ_ (Hom i j)).inv ≫ Hom i j ◁ id j ≫ comp i j j = 𝟙 (Hom i j) := by + rw [Iso.inv_comp_eq] + ext + exact Functor.ext (fun _ ↦ by simp) + +@[simp] +lemma assoc (i j k l : SimplicialThickening J) : + (α_ (Hom i j) (Hom j k) (Hom k l)).inv ≫ comp i j k ▷ Hom k l ≫ comp i k l = + Hom i j ◁ comp j k l ≫ comp i j l := by + ext + exact Functor.ext (fun _ ↦ by simp) + +end SimplicialCategory + +open SimplicialThickening.SimplicialCategory + +noncomputable instance (J : Type*) [LinearOrder J] : + SimplicialCategory (SimplicialThickening J) where + Hom := Hom + id := id + comp := comp + homEquiv {i j} := (nerveEquiv _).symm.trans (SSet.unitHomEquiv _).symm + +/-- Auxiliary definition for `SimplicialThickening.functorMap` -/ +def orderHom {J K : Type*} [LinearOrder J] [LinearOrder K] (f : J →o K) : + SimplicialThickening J →o SimplicialThickening K := f + +/-- Auxiliary definition for `SimplicialThickening.functor` -/ +noncomputable abbrev functorMap {J K : Type u} [LinearOrder J] [LinearOrder K] + (f : J →o K) (i j : SimplicialThickening J) : (i ⟶ j) ⥤ ((orderHom f i) ⟶ (orderHom f j)) where + obj I := ⟨f '' I.I, Set.mem_image_of_mem f I.left, Set.mem_image_of_mem f I.right, + by rintro _ ⟨k, hk, rfl⟩; exact f.monotone (I.left_le k hk), + by rintro _ ⟨k, hk, rfl⟩; exact f.monotone (I.le_right k hk)⟩ + map f := ⟨⟨Set.image_subset _ f.1.1⟩⟩ + +/-- +The simplicial thickening defines a functor from the category of linear orders to the category of +simplicial categories +-/ +@[simps] +noncomputable def functor {J K : Type u} [LinearOrder J] [LinearOrder K] + (f : J →o K) : EnrichedFunctor SSet (SimplicialThickening J) (SimplicialThickening K) where + obj := f + map i j := nerveMap ((functorMap f i j)) + map_id i := by + ext + simp only [eId, EnrichedCategory.id] + exact Functor.ext (by aesop_cat) + map_comp i j k := by + ext + simp only [eComp, EnrichedCategory.comp] + exact Functor.ext (by aesop_cat) + +lemma functor_id (J : Type u) [LinearOrder J] : + (functor (OrderHom.id (α := J))) = EnrichedFunctor.id _ _ := by + refine EnrichedFunctor.ext _ (fun _ ↦ rfl) fun i j ↦ ?_ + ext + exact Functor.ext (by aesop_cat) + +lemma functor_comp {J K L : Type u} [LinearOrder J] [LinearOrder K] + [LinearOrder L] (f : J →o K) (g : K →o L) : + functor (g.comp f) = + (functor f).comp _ (functor g) := by + refine EnrichedFunctor.ext _ (fun _ ↦ rfl) fun i j ↦ ?_ + ext + exact Functor.ext (by aesop_cat) + +end SimplicialThickening + +/-- +The simplicial nerve of a simplicial category `C` is defined as the simplicial set whose +`n`-simplices are given by the set of simplicial functors from the simplicial thickening of +the linear order `Fin (n + 1)` to `C` +-/ +noncomputable def SimplicialNerve (C : Type u) [Category.{v} C] [SimplicialCategory C] : + SSet.{max u v} where + obj n := EnrichedFunctor SSet (SimplicialThickening (ULift (Fin (n.unop.len + 1)))) C + map f := (SimplicialThickening.functor f.unop.toOrderHom.uliftMap).comp (E := C) SSet + map_id i := by + change EnrichedFunctor.comp SSet (SimplicialThickening.functor (OrderHom.id)) = _ + rw [SimplicialThickening.functor_id] + rfl + map_comp f g := by + change EnrichedFunctor.comp SSet (SimplicialThickening.functor + (f.unop.toOrderHom.uliftMap.comp g.unop.toOrderHom.uliftMap)) = _ + rw [SimplicialThickening.functor_comp] + rfl + +end SimplicialNerve + +end CategoryTheory diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean index be2e7617b9773..bc6345190d454 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean @@ -61,6 +61,10 @@ instance hasColimits : HasColimits SSet := by lemma hom_ext {X Y : SSet} {f g : X ⟶ Y} (w : ∀ n, f.app n = g.app n) : f = g := SimplicialObject.hom_ext _ _ w +@[simp] +lemma comp_app {X Y Z : SSet} (f : X ⟶ Y) (g : Y ⟶ Z) (n : SimplexCategoryᵒᵖ) : + (f ≫ g).app n = f.app n ≫ g.app n := NatTrans.comp_app _ _ _ + /-- The ulift functor `SSet.{u} ⥤ SSet.{max u v}` on simplicial sets. -/ def uliftFunctor : SSet.{u} ⥤ SSet.{max u v} := (SimplicialObject.whiskering _ _).obj CategoryTheory.uliftFunctor.{v, u} diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Nerve.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Nerve.lean index a24814b5a2a14..0f673d49480b6 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Nerve.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Nerve.lean @@ -35,11 +35,23 @@ def nerve (C : Type u) [Category.{v} C] : SSet.{max u v} where instance {C : Type*} [Category C] {Δ : SimplexCategoryᵒᵖ} : Category ((nerve C).obj Δ) := (inferInstance : Category (ComposableArrows C (Δ.unop.len))) +/-- Given a functor `C ⥤ D`, we obtain a morphism `nerve C ⟶ nerve D` of simplicial sets. -/ +@[simps] +def nerveMap {C D : Type u} [Category.{v} C] [Category.{v} D] (F : C ⥤ D) : nerve C ⟶ nerve D := + { app := fun _ => (F.mapComposableArrows _).obj } + /-- The nerve of a category, as a functor `Cat ⥤ SSet` -/ @[simps] def nerveFunctor : Cat.{v, u} ⥤ SSet where obj C := nerve C - map F := { app := fun _ => (F.mapComposableArrows _).obj } + map F := nerveMap F + +/-- The 0-simplices of the nerve of a category are equivalent to the objects of the category. -/ +def nerveEquiv (C : Type u) [Category.{v} C] : nerve C _[0] ≃ C where + toFun f := f.obj ⟨0, by omega⟩ + invFun f := (Functor.const _).obj f + left_inv f := ComposableArrows.ext₀ rfl + right_inv f := rfl namespace Nerve diff --git a/Mathlib/CategoryTheory/Enriched/Basic.lean b/Mathlib/CategoryTheory/Enriched/Basic.lean index 239f6399ae5be..bd179b0fdb051 100644 --- a/Mathlib/CategoryTheory/Enriched/Basic.lean +++ b/Mathlib/CategoryTheory/Enriched/Basic.lean @@ -297,6 +297,16 @@ def EnrichedFunctor.comp {C : Type u₁} {D : Type u₂} {E : Type u₃} [Enrich obj X := G.obj (F.obj X) map _ _ := F.map _ _ ≫ G.map _ _ +lemma EnrichedFunctor.ext {C : Type u₁} {D : Type u₂} [EnrichedCategory V C] + [EnrichedCategory V D] {F G : EnrichedFunctor V C D} (h_obj : ∀ X, F.obj X = G.obj X) + (h_map : ∀ (X Y : C), F.map X Y ≫ eqToHom (by rw [h_obj, h_obj]) = G.map X Y) : F = G := by + match F, G with + | mk F_obj F_map _ _, mk G_obj G_map _ _ => + obtain rfl : F_obj = G_obj := funext fun X ↦ h_obj X + congr + ext X Y + simpa using h_map X Y + section variable {W : Type (v + 1)} [Category.{v} W] [MonoidalCategory W] diff --git a/Mathlib/Order/Hom/Basic.lean b/Mathlib/Order/Hom/Basic.lean index ef9ebbb809b7f..4e602339f6803 100644 --- a/Mathlib/Order/Hom/Basic.lean +++ b/Mathlib/Order/Hom/Basic.lean @@ -521,6 +521,12 @@ protected def withBotMap (f : α →o β) : WithBot α →o WithBot β := protected def withTopMap (f : α →o β) : WithTop α →o WithTop β := ⟨WithTop.map f, f.mono.withTop_map⟩ +/-- Lift an order homomorphism `f : α →o β` to an order homomorphism `ULift α →o ULift β` in a +higher universe. -/ +@[simps!] +def uliftMap (f : α →o β) : ULift α →o ULift β := + ⟨fun i => ⟨f i.down⟩, fun _ _ h ↦ f.monotone h⟩ + end OrderHom /-- Embeddings of partial orders that preserve `<` also preserve `≤`. -/ diff --git a/docs/references.bib b/docs/references.bib index 69f11f4a17b95..d397a999563cc 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -2557,6 +2557,13 @@ @InProceedings{ lewis2019 keywords = {Hensel's lemma, Lean, formal proof, p-adic} } +@Book{ LurieHTT, + title = {Higher Topos Theory}, + author = {Jacob Lurie}, + url = {https://www.math.ias.edu/~lurie/papers/HTT.pdf}, + year = {2009} +} + @Book{ LurieSAG, title = {Spectral Algebraic Geometry}, author = {Jacob Lurie}, From 407823ad8837bb81e4d314806c652b2131496a91 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Thu, 12 Dec 2024 14:48:32 +0000 Subject: [PATCH 679/829] chore: update Mathlib dependencies 2024-12-12 (#19917) 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 df098d92919ba..09adba8286a84 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "74dffd1a83cdd2969a31c9892b0517e7c6f50668", + "rev": "9e583efcea920afa13ee2a53069821a2297a94c0", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From 808600f8f099dc32f03fd226140ffcbddbddf71b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Thu, 12 Dec 2024 16:03:40 +0000 Subject: [PATCH 680/829] chore: delete `nonempty_toType_aleph` (#19919) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This instance was a byproduct of defining [`MeasurableSpace.generateMeasurableRec`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/MeasureTheory/MeasurableSpace/Card.html#MeasurableSpace.generateMeasurableRec) in terms of the very ugly `(ℵ_ o).ord.toType`. Now that we properly define it in terms of an ordinal, we have no need for this (and frankly we should not encourage it either). --- Mathlib/SetTheory/Cardinal/Aleph.lean | 4 ---- 1 file changed, 4 deletions(-) diff --git a/Mathlib/SetTheory/Cardinal/Aleph.lean b/Mathlib/SetTheory/Cardinal/Aleph.lean index 52eabc3d86978..db72a7115f40d 100644 --- a/Mathlib/SetTheory/Cardinal/Aleph.lean +++ b/Mathlib/SetTheory/Cardinal/Aleph.lean @@ -427,10 +427,6 @@ theorem aleph_toNat (o : Ordinal) : toNat (ℵ_ o) = 0 := theorem aleph_toENat (o : Ordinal) : toENat (ℵ_ o) = ⊤ := (toENat_eq_top.2 (aleph0_le_aleph o)) -instance nonempty_toType_aleph (o : Ordinal) : Nonempty (ℵ_ o).ord.toType := by - rw [toType_nonempty_iff_ne_zero, ← ord_zero] - exact fun h => (ord_injective h).not_gt (aleph_pos o) - theorem isLimit_omega (o : Ordinal) : Ordinal.IsLimit (ω_ o) := by rw [← ord_aleph] exact isLimit_ord (aleph0_le_aleph _) From fdbd2e1a2b466abf820676ba068d139f2e5d2635 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 12 Dec 2024 18:01:33 +0000 Subject: [PATCH 681/829] feat(GiryMonad): add `Measurable.measure_of_isPiSystem` (#19908) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit ... and use it to golf some `Measurable (_ : _ → Measure _)` theorems. --- Mathlib/MeasureTheory/Measure/GiryMonad.lean | 18 ++++++++++ .../Kernel/Disintegration/CDFToKernel.lean | 28 --------------- .../Kernel/Disintegration/CondCDF.lean | 26 +++----------- .../Disintegration/MeasurableStieltjes.lean | 34 ++++++++----------- 4 files changed, 37 insertions(+), 69 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/GiryMonad.lean b/Mathlib/MeasureTheory/Measure/GiryMonad.lean index 4b954b6e64486..fd5a97331a2ec 100644 --- a/Mathlib/MeasureTheory/Measure/GiryMonad.lean +++ b/Mathlib/MeasureTheory/Measure/GiryMonad.lean @@ -64,6 +64,24 @@ theorem measurable_measure {μ : α → Measure β} : Measurable μ ↔ ∀ (s : Set β), MeasurableSet s → Measurable fun b => μ b s := ⟨fun hμ _s hs => (measurable_coe hs).comp hμ, measurable_of_measurable_coe μ⟩ +theorem _root_.Measurable.measure_of_isPiSystem {μ : α → Measure β} [∀ a, IsFiniteMeasure (μ a)] + {S : Set (Set β)} (hgen : ‹MeasurableSpace β› = .generateFrom S) (hpi : IsPiSystem S) + (h_basic : ∀ s ∈ S, Measurable fun a ↦ μ a s) (h_univ : Measurable fun a ↦ μ a univ) : + Measurable μ := by + rw [measurable_measure] + refine MeasurableSpace.induction_on_inter hgen hpi (by simp) h_basic ?_ ?_ + · intro s hsm ihs + simp only [measure_compl hsm (measure_ne_top _ _)] + exact h_univ.sub ihs + · intro f hfd hfm ihf + simpa only [measure_iUnion hfd hfm] using .ennreal_tsum ihf + +theorem _root_.Measurable.measure_of_isPiSystem_of_isProbabilityMeasure {μ : α → Measure β} + [∀ a, IsProbabilityMeasure (μ a)] + {S : Set (Set β)} (hgen : ‹MeasurableSpace β› = .generateFrom S) (hpi : IsPiSystem S) + (h_basic : ∀ s ∈ S, Measurable fun a ↦ μ a s) : Measurable μ := + .measure_of_isPiSystem hgen hpi h_basic <| by simp + theorem measurable_map (f : α → β) (hf : Measurable f) : Measurable fun μ : Measure α => map f μ := by refine measurable_of_measurable_coe _ fun s hs => ?_ diff --git a/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean index a8558585a52da..b1fdd2d1f0025 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean @@ -471,34 +471,6 @@ section ToKernel variable {_ : MeasurableSpace β} {f : α × β → StieltjesFunction} {κ : Kernel α (β × ℝ)} {ν : Kernel α β} -/-- A measurable function `α → StieltjesFunction` with limits 0 at -∞ and 1 at +∞ gives a measurable -function `α → Measure ℝ` by taking `StieltjesFunction.measure` at each point. -/ -lemma StieltjesFunction.measurable_measure {f : α → StieltjesFunction} - (hf : ∀ q, Measurable fun a ↦ f a q) - (hf_bot : ∀ a, Tendsto (f a) atBot (𝓝 0)) - (hf_top : ∀ a, Tendsto (f a) atTop (𝓝 1)) : - Measurable fun a ↦ (f a).measure := by - refine Measure.measurable_measure.mpr fun s hs ↦ ?_ - have : ∀ a, IsProbabilityMeasure (f a).measure := - fun a ↦ (f a).isProbabilityMeasure (hf_bot a) (hf_top a) - refine MeasurableSpace.induction_on_inter - (C := fun s ↦ Measurable fun b ↦ StieltjesFunction.measure (f b) s) - (borel_eq_generateFrom_Iic ℝ) isPiSystem_Iic ?_ ?_ ?_ ?_ hs - · simp only [measure_empty, measurable_const] - · rintro S ⟨u, rfl⟩ - simp_rw [StieltjesFunction.measure_Iic (f _) (hf_bot _), sub_zero] - exact (hf _).ennreal_ofReal - · intro t ht ht_cd_meas - have : (fun a ↦ (f a).measure tᶜ) = (fun a ↦ (f a).measure univ) - fun a ↦ (f a).measure t := by - ext1 a - rw [measure_compl ht, Pi.sub_apply] - exact measure_ne_top _ _ - simp_rw [this, measure_univ] - exact Measurable.sub measurable_const ht_cd_meas - · intro f hf_disj hf_meas hf_cd_meas - simp_rw [measure_iUnion hf_disj hf_meas] - exact Measurable.ennreal_tsum hf_cd_meas - /-- A function `f : α × β → StieltjesFunction` with the property `IsCondKernelCDF f κ ν` gives a Markov kernel from `α × β` to `ℝ`, by taking for each `p : α × β` the measure defined by `f p`. -/ noncomputable diff --git a/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean b/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean index 9d6ed97e8708e..7a63f0c847a34 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean @@ -335,28 +335,10 @@ instance instIsProbabilityMeasureCondCDF (ρ : Measure (α × ℝ)) (a : α) : /-- The function `a ↦ (condCDF ρ a).measure` is measurable. -/ theorem measurable_measure_condCDF (ρ : Measure (α × ℝ)) : - Measurable fun a => (condCDF ρ a).measure := by - rw [Measure.measurable_measure] - refine fun s hs => ?_ - -- Porting note: supplied `C` - refine MeasurableSpace.induction_on_inter - (C := fun s => Measurable fun b ↦ StieltjesFunction.measure (condCDF ρ b) s) - (borel_eq_generateFrom_Iic ℝ) isPiSystem_Iic ?_ ?_ ?_ ?_ hs - · simp only [measure_empty, measurable_const] - · rintro S ⟨u, rfl⟩ - simp_rw [measure_condCDF_Iic ρ _ u] - exact (measurable_condCDF ρ u).ennreal_ofReal - · intro t ht ht_cd_meas - have : - (fun a => (condCDF ρ a).measure tᶜ) = - (fun a => (condCDF ρ a).measure univ) - fun a => (condCDF ρ a).measure t := by - ext1 a - rw [measure_compl ht (measure_ne_top (condCDF ρ a).measure _), Pi.sub_apply] - simp_rw [this, measure_condCDF_univ ρ] - exact Measurable.sub measurable_const ht_cd_meas - · intro f hf_disj hf_meas hf_cd_meas - simp_rw [measure_iUnion hf_disj hf_meas] - exact Measurable.ennreal_tsum hf_cd_meas + Measurable fun a => (condCDF ρ a).measure := + .measure_of_isPiSystem_of_isProbabilityMeasure (borel_eq_generateFrom_Iic ℝ) isPiSystem_Iic <| by + simp_rw [forall_mem_range, measure_condCDF_Iic] + exact fun u ↦ (measurable_condCDF ρ u).ennreal_ofReal end Measure diff --git a/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean index 251bef27aec44..a3cea10685cc2 100644 --- a/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean +++ b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean @@ -45,6 +45,19 @@ open MeasureTheory Set Filter TopologicalSpace open scoped NNReal ENNReal MeasureTheory Topology +/-- A measurable function `α → StieltjesFunction` with limits 0 at -∞ and 1 at +∞ gives a measurable +function `α → Measure ℝ` by taking `StieltjesFunction.measure` at each point. -/ +lemma StieltjesFunction.measurable_measure {α : Type*} {_ : MeasurableSpace α} + {f : α → StieltjesFunction} (hf : ∀ q, Measurable fun a ↦ f a q) + (hf_bot : ∀ a, Tendsto (f a) atBot (𝓝 0)) + (hf_top : ∀ a, Tendsto (f a) atTop (𝓝 1)) : + Measurable fun a ↦ (f a).measure := + have : ∀ a, IsProbabilityMeasure (f a).measure := + fun a ↦ (f a).isProbabilityMeasure (hf_bot a) (hf_top a) + .measure_of_isPiSystem_of_isProbabilityMeasure (borel_eq_generateFrom_Iic ℝ) isPiSystem_Iic <| by + simp_rw [forall_mem_range, StieltjesFunction.measure_Iic (f _) (hf_bot _), sub_zero] + exact fun _ ↦ (hf _).ennreal_ofReal + namespace ProbabilityTheory variable {α : Type*} @@ -412,25 +425,8 @@ instance IsMeasurableRatCDF.instIsProbabilityMeasure_stieltjesFunction (a : α) lemma IsMeasurableRatCDF.measurable_measure_stieltjesFunction : Measurable fun a ↦ (hf.stieltjesFunction a).measure := by - rw [Measure.measurable_measure] - refine fun s hs ↦ MeasurableSpace.induction_on_inter - (C := fun s ↦ Measurable fun b ↦ StieltjesFunction.measure (hf.stieltjesFunction b) s) - (borel_eq_generateFrom_Iic ℝ) isPiSystem_Iic ?_ ?_ ?_ ?_ hs - · simp only [measure_empty, measurable_const] - · rintro S ⟨u, rfl⟩ - simp_rw [measure_stieltjesFunction_Iic hf _ u] - exact (measurable_stieltjesFunction hf u).ennreal_ofReal - · intro t ht ht_cd_meas - have : (fun a ↦ (hf.stieltjesFunction a).measure tᶜ) = - (fun a ↦ (hf.stieltjesFunction a).measure univ) - - fun a ↦ (hf.stieltjesFunction a).measure t := by - ext1 a - rw [measure_compl ht (measure_ne_top (hf.stieltjesFunction a).measure _), Pi.sub_apply] - simp_rw [this, measure_stieltjesFunction_univ hf] - exact Measurable.sub measurable_const ht_cd_meas - · intro f hf_disj hf_meas hf_cd_meas - simp_rw [measure_iUnion hf_disj hf_meas] - exact Measurable.ennreal_tsum hf_cd_meas + apply_rules [StieltjesFunction.measurable_measure, measurable_stieltjesFunction, + tendsto_stieltjesFunction_atBot, tendsto_stieltjesFunction_atTop] end Measure From 54c954e5095b14a0fca38552f6d04ac165e59d11 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 12 Dec 2024 18:16:22 +0000 Subject: [PATCH 682/829] refactor(fderiv): redefine `fderivWithin` and `fderiv` (#19694) Prefer `0` answer whenever it's possible. This way, `fderivWithin_const` is true without extra assumptions. --- Mathlib/Analysis/Calculus/ContDiff/Basic.lean | 56 +++++----- .../Analysis/Calculus/ContDiff/Bounds.lean | 2 +- Mathlib/Analysis/Calculus/Deriv/Basic.lean | 5 +- Mathlib/Analysis/Calculus/Deriv/Slope.lean | 2 +- Mathlib/Analysis/Calculus/FDeriv/Add.lean | 26 +---- Mathlib/Analysis/Calculus/FDeriv/Basic.lean | 105 ++++++++---------- 6 files changed, 77 insertions(+), 119 deletions(-) diff --git a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean index 7b0fc7934c102..2cab2349803e7 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean @@ -54,21 +54,26 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] /-! ### Constants -/ +theorem iteratedFDerivWithin_succ_const (n : ℕ) (c : F) : + iteratedFDerivWithin 𝕜 (n + 1) (fun _ : E ↦ c) s = 0 := by + induction n with + | zero => + ext1 + simp [iteratedFDerivWithin_succ_eq_comp_left, iteratedFDerivWithin_zero_eq_comp, comp_def] + | succ n IH => + rw [iteratedFDerivWithin_succ_eq_comp_left, IH] + simp only [Pi.zero_def, comp_def, fderivWithin_const, map_zero] + @[simp] -theorem iteratedFDerivWithin_zero_fun (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {i : ℕ} : - iteratedFDerivWithin 𝕜 i (fun _ : E ↦ (0 : F)) s x = 0 := by - induction i generalizing x with +theorem iteratedFDerivWithin_zero_fun {i : ℕ} : + iteratedFDerivWithin 𝕜 i (fun _ : E ↦ (0 : F)) s = 0 := by + cases i with | zero => ext; simp - | succ i IH => - ext m - rw [iteratedFDerivWithin_succ_apply_left, fderivWithin_congr (fun _ ↦ IH) (IH hx)] - rw [fderivWithin_const_apply _ (hs x hx)] - rfl + | succ i => apply iteratedFDerivWithin_succ_const @[simp] theorem iteratedFDeriv_zero_fun {n : ℕ} : (iteratedFDeriv 𝕜 n fun _ : E ↦ (0 : F)) = 0 := - funext fun x ↦ by simpa [← iteratedFDerivWithin_univ] using - iteratedFDerivWithin_zero_fun uniqueDiffOn_univ (mem_univ x) + funext fun x ↦ by simp only [← iteratedFDerivWithin_univ, iteratedFDerivWithin_zero_fun] theorem contDiff_zero_fun : ContDiff 𝕜 n fun _ : E => (0 : F) := analyticOnNhd_const.contDiff @@ -103,30 +108,19 @@ theorem contDiffWithinAt_of_subsingleton [Subsingleton F] : ContDiffWithinAt theorem contDiffOn_of_subsingleton [Subsingleton F] : ContDiffOn 𝕜 n f s := by rw [Subsingleton.elim f fun _ => 0]; exact contDiffOn_const -theorem iteratedFDerivWithin_succ_const (n : ℕ) (c : F) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) : - iteratedFDerivWithin 𝕜 (n + 1) (fun _ : E ↦ c) s x = 0 := by - ext m - rw [iteratedFDerivWithin_succ_apply_right hs hx] - rw [iteratedFDerivWithin_congr (fun y hy ↦ fderivWithin_const_apply c (hs y hy)) hx] - rw [iteratedFDerivWithin_zero_fun hs hx] - simp [ContinuousMultilinearMap.zero_apply (R := 𝕜)] - -theorem iteratedFDeriv_succ_const (n : ℕ) (c : F) : - (iteratedFDeriv 𝕜 (n + 1) fun _ : E ↦ c) = 0 := - funext fun x ↦ by simpa [← iteratedFDerivWithin_univ] using - iteratedFDerivWithin_succ_const n c uniqueDiffOn_univ (mem_univ x) - -theorem iteratedFDerivWithin_const_of_ne {n : ℕ} (hn : n ≠ 0) (c : F) - (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) : - iteratedFDerivWithin 𝕜 n (fun _ : E ↦ c) s x = 0 := by +theorem iteratedFDerivWithin_const_of_ne {n : ℕ} (hn : n ≠ 0) (c : F) (s : Set E) : + iteratedFDerivWithin 𝕜 n (fun _ : E ↦ c) s = 0 := by cases n with | zero => contradiction - | succ n => exact iteratedFDerivWithin_succ_const n c hs hx + | succ n => exact iteratedFDerivWithin_succ_const n c theorem iteratedFDeriv_const_of_ne {n : ℕ} (hn : n ≠ 0) (c : F) : - (iteratedFDeriv 𝕜 n fun _ : E ↦ c) = 0 := - funext fun x ↦ by simpa [← iteratedFDerivWithin_univ] using - iteratedFDerivWithin_const_of_ne hn c uniqueDiffOn_univ (mem_univ x) + (iteratedFDeriv 𝕜 n fun _ : E ↦ c) = 0 := by + simp only [← iteratedFDerivWithin_univ, iteratedFDerivWithin_const_of_ne hn] + +theorem iteratedFDeriv_succ_const (n : ℕ) (c : F) : + (iteratedFDeriv 𝕜 (n + 1) fun _ : E ↦ c) = 0 := + iteratedFDeriv_const_of_ne (by simp) _ theorem contDiffWithinAt_singleton : ContDiffWithinAt 𝕜 n f {x} x := (contDiffWithinAt_const (c := f x)).congr (by simp) rfl @@ -946,7 +940,7 @@ theorem iteratedFDerivWithin_clm_apply_const_apply rw [fderivWithin_congr' (fun x hx ↦ ih hi.le hx) hx] rw [fderivWithin_clm_apply (hs x hx) (h_deriv.continuousMultilinear_apply_const _ x hx) (differentiableWithinAt_const u)] - rw [fderivWithin_const_apply _ (hs x hx)] + rw [fderivWithin_const_apply] simp only [ContinuousLinearMap.flip_apply, ContinuousLinearMap.comp_zero, zero_add] rw [fderivWithin_continuousMultilinear_apply_const_apply (hs x hx) (h_deriv x hx)] diff --git a/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean b/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean index 008902d921c30..c869df9c43a51 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean @@ -296,7 +296,7 @@ theorem norm_iteratedFDerivWithin_prod_le [DecidableEq ι] [NormOneClass A'] {u | empty => cases n with | zero => simp [Sym.eq_nil_of_card_zero] - | succ n => simp [iteratedFDerivWithin_succ_const _ _ hs hx] + | succ n => simp [iteratedFDerivWithin_succ_const] | @insert i u hi IH => conv => lhs; simp only [Finset.prod_insert hi] simp only [Finset.mem_insert, forall_eq_or_imp] at hf diff --git a/Mathlib/Analysis/Calculus/Deriv/Basic.lean b/Mathlib/Analysis/Calculus/Deriv/Basic.lean index 2980e8e408490..cd89729cd9472 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Basic.lean @@ -630,8 +630,9 @@ theorem deriv_const : deriv (fun _ => c) x = 0 := theorem deriv_const' : (deriv fun _ : 𝕜 => c) = fun _ => 0 := funext fun x => deriv_const x c -theorem derivWithin_const (hxs : UniqueDiffWithinAt 𝕜 s x) : derivWithin (fun _ => c) s x = 0 := - (hasDerivWithinAt_const _ _ _).derivWithin hxs +@[simp] +theorem derivWithin_const : derivWithin (fun _ => c) s = 0 := by + ext; simp [derivWithin] end Const diff --git a/Mathlib/Analysis/Calculus/Deriv/Slope.lean b/Mathlib/Analysis/Calculus/Deriv/Slope.lean index b50f39e69ad53..20e80a953290f 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Slope.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Slope.lean @@ -92,7 +92,7 @@ theorem range_derivWithin_subset_closure_span_image range (derivWithin f s) ⊆ closure (Submodule.span 𝕜 (f '' t)) := by rintro - ⟨x, rfl⟩ rcases eq_or_neBot (𝓝[s \ {x}] x) with H|H - · simpa [derivWithin, fderivWithin, H] using subset_closure (zero_mem _) + · simpa [derivWithin_zero_of_isolated H] using subset_closure (zero_mem _) by_cases H' : DifferentiableWithinAt 𝕜 f s x; swap · rw [derivWithin_zero_of_not_differentiableWithinAt H'] exact subset_closure (zero_mem _) diff --git a/Mathlib/Analysis/Calculus/FDeriv/Add.lean b/Mathlib/Analysis/Calculus/FDeriv/Add.lean index e6fb07b97cd41..fd877bcc04666 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Add.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Add.lean @@ -739,31 +739,7 @@ theorem differentiableWithinAt_comp_add_right (a : E) : theorem fderivWithin_comp_add_right (a : E) : fderivWithin 𝕜 (fun x ↦ f (x + a)) s x = fderivWithin 𝕜 f (a +ᵥ s) (x + a) := by - simp only [fderivWithin, hasFDerivWithinAt_comp_add_right] - by_cases h : 𝓝[s \ {x}] x = ⊥ - · have h' : 𝓝[(a +ᵥ s) \ {x + a}] (x + a) = ⊥ := by - let e := Homeomorph.addRight a - have : 𝓝[(a +ᵥ s) \ {x + a}] (x + a) = map e (𝓝[s \ {x}] x) := by - rw [e.isEmbedding.map_nhdsWithin_eq] - congr - ext y - rw [← e.preimage_symm] - simp [e, Homeomorph.addRight, Set.mem_vadd_set_iff_neg_vadd_mem, add_comm] - rw [this, h, Filter.map_bot] - simp [h, h'] - · have h' : 𝓝[(a +ᵥ s) \ {x + a}] (x + a) ≠ ⊥ := by - intro h'' - let e := Homeomorph.addRight (-a) - have : 𝓝[s \ {x}] x = map e (𝓝[(a +ᵥ s) \ {x + a}] (x + a)) := by - rw [e.isEmbedding.map_nhdsWithin_eq] - congr - · simp [e] - ext y - rw [← e.preimage_symm] - simp [e, Homeomorph.addRight, Set.mem_vadd_set_iff_neg_vadd_mem, add_comm] - apply h - rw [this, h'', Filter.map_bot] - simp [h, h'] + simp only [fderivWithin, hasFDerivWithinAt_comp_add_right, DifferentiableWithinAt] theorem hasFDerivWithinAt_comp_add_left (a : E) : HasFDerivWithinAt (fun x ↦ f (a + x)) f' s x ↔ HasFDerivWithinAt f f' (a +ᵥ s) (a + x) := by diff --git a/Mathlib/Analysis/Calculus/FDeriv/Basic.lean b/Mathlib/Analysis/Calculus/FDeriv/Basic.lean index c381a83392646..aa719cbb852ef 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Basic.lean @@ -171,16 +171,18 @@ def DifferentiableAt (f : E → F) (x : E) := ∃ f' : E →L[𝕜] F, HasFDerivAt f f' x /-- If `f` has a derivative at `x` within `s`, then `fderivWithin 𝕜 f s x` is such a derivative. -Otherwise, it is set to `0`. If `x` is isolated in `s`, we take the derivative within `s` to -be zero for convenience. -/ +Otherwise, it is set to `0`. We also set it to be zero, if zero is one of possible derivatives. -/ irreducible_def fderivWithin (f : E → F) (s : Set E) (x : E) : E →L[𝕜] F := - if 𝓝[s \ {x}] x = ⊥ then 0 else - if h : ∃ f', HasFDerivWithinAt f f' s x then Classical.choose h else 0 + if HasFDerivWithinAt f (0 : E →L[𝕜] F) s x + then 0 + else if h : DifferentiableWithinAt 𝕜 f s x + then Classical.choose h + else 0 /-- If `f` has a derivative at `x`, then `fderiv 𝕜 f x` is such a derivative. Otherwise, it is set to `0`. -/ irreducible_def fderiv (f : E → F) (x : E) : E →L[𝕜] F := - if h : ∃ f', HasFDerivAt f f' x then Classical.choose h else 0 + fderivWithin 𝕜 f univ x /-- `DifferentiableOn 𝕜 f s` means that `f` is differentiable within `s` at any point of `s`. -/ @[fun_prop] @@ -199,23 +201,14 @@ variable {x : E} variable {s t : Set E} variable {L L₁ L₂ : Filter E} -theorem fderivWithin_zero_of_isolated (h : 𝓝[s \ {x}] x = ⊥) : fderivWithin 𝕜 f s x = 0 := by - rw [fderivWithin, if_pos h] - -theorem fderivWithin_zero_of_nmem_closure (h : x ∉ closure s) : fderivWithin 𝕜 f s x = 0 := by - apply fderivWithin_zero_of_isolated - simp only [mem_closure_iff_nhdsWithin_neBot, neBot_iff, Ne, Classical.not_not] at h - rw [eq_bot_iff, ← h] - exact nhdsWithin_mono _ diff_subset - theorem fderivWithin_zero_of_not_differentiableWithinAt (h : ¬DifferentiableWithinAt 𝕜 f s x) : fderivWithin 𝕜 f s x = 0 := by - have : ¬∃ f', HasFDerivWithinAt f f' s x := h - simp [fderivWithin, this] + simp [fderivWithin, h] -theorem fderiv_zero_of_not_differentiableAt (h : ¬DifferentiableAt 𝕜 f x) : fderiv 𝕜 f x = 0 := by - have : ¬∃ f', HasFDerivAt f f' x := h - simp [fderiv, this] +@[simp] +theorem fderivWithin_univ : fderivWithin 𝕜 f univ = fderiv 𝕜 f := by + ext + rw [fderiv] section DerivativeUniqueness @@ -375,6 +368,14 @@ theorem hasFDerivWithinAt_univ : HasFDerivWithinAt f f' univ x ↔ HasFDerivAt f alias ⟨HasFDerivWithinAt.hasFDerivAt_of_univ, _⟩ := hasFDerivWithinAt_univ +theorem differentiableWithinAt_univ : + DifferentiableWithinAt 𝕜 f univ x ↔ DifferentiableAt 𝕜 f x := by + simp only [DifferentiableWithinAt, hasFDerivWithinAt_univ, DifferentiableAt] + +theorem fderiv_zero_of_not_differentiableAt (h : ¬DifferentiableAt 𝕜 f x) : fderiv 𝕜 f x = 0 := by + rw [fderiv, fderivWithin_zero_of_not_differentiableWithinAt] + rwa [differentiableWithinAt_univ] + theorem hasFDerivWithinAt_of_mem_nhds (h : s ∈ 𝓝 x) : HasFDerivWithinAt f f' s x ↔ HasFDerivAt f f' x := by rw [HasFDerivAt, HasFDerivWithinAt, nhdsWithin_eq_nhds.mpr h] @@ -486,19 +487,23 @@ theorem hasFDerivWithinAt_of_nmem_closure (h : x ∉ closure s) : HasFDerivWithi .of_nhdsWithin_eq_bot <| eq_bot_mono (nhdsWithin_mono _ diff_subset) <| by rwa [mem_closure_iff_nhdsWithin_neBot, not_neBot] at h +theorem fderivWithin_zero_of_isolated (h : 𝓝[s \ {x}] x = ⊥) : fderivWithin 𝕜 f s x = 0 := by + rw [fderivWithin, if_pos (.of_nhdsWithin_eq_bot h)] + +theorem fderivWithin_zero_of_nmem_closure (h : x ∉ closure s) : fderivWithin 𝕜 f s x = 0 := by + rw [fderivWithin, if_pos (hasFDerivWithinAt_of_nmem_closure h)] + theorem DifferentiableWithinAt.hasFDerivWithinAt (h : DifferentiableWithinAt 𝕜 f s x) : HasFDerivWithinAt f (fderivWithin 𝕜 f s x) s x := by - by_cases H : 𝓝[s \ {x}] x = ⊥ - · exact .of_nhdsWithin_eq_bot H - · unfold DifferentiableWithinAt at h - rw [fderivWithin, if_neg H, dif_pos h] - exact Classical.choose_spec h + simp only [fderivWithin, dif_pos h] + split_ifs with h₀ + exacts [h₀, Classical.choose_spec h] theorem DifferentiableAt.hasFDerivAt (h : DifferentiableAt 𝕜 f x) : HasFDerivAt f (fderiv 𝕜 f x) x := by - dsimp only [DifferentiableAt] at h - rw [fderiv, dif_pos h] - exact Classical.choose_spec h + rw [fderiv, ← hasFDerivWithinAt_univ] + rw [← differentiableWithinAt_univ] at h + exact h.hasFDerivWithinAt theorem DifferentiableOn.hasFDerivAt (h : DifferentiableOn 𝕜 f s) (hs : s ∈ 𝓝 x) : HasFDerivAt f (fderiv 𝕜 f x) x := @@ -575,10 +580,6 @@ theorem differentiableWithinAt_congr_nhds {t : Set E} (hst : 𝓝[s] x = 𝓝[t] DifferentiableWithinAt 𝕜 f s x ↔ DifferentiableWithinAt 𝕜 f t x := ⟨fun h => h.congr_nhds hst, fun h => h.congr_nhds hst.symm⟩ -theorem differentiableWithinAt_univ : - DifferentiableWithinAt 𝕜 f univ x ↔ DifferentiableAt 𝕜 f x := by - simp only [DifferentiableWithinAt, hasFDerivWithinAt_univ, DifferentiableAt] - theorem differentiableWithinAt_inter (ht : t ∈ 𝓝 x) : DifferentiableWithinAt 𝕜 f (s ∩ t) x ↔ DifferentiableWithinAt 𝕜 f s x := by simp only [DifferentiableWithinAt, hasFDerivWithinAt_inter ht] @@ -647,19 +648,7 @@ theorem fderivWithin_subset (st : s ⊆ t) (ht : UniqueDiffWithinAt 𝕜 s x) fderivWithin_of_mem_nhdsWithin (nhdsWithin_mono _ st self_mem_nhdsWithin) ht h theorem fderivWithin_inter (ht : t ∈ 𝓝 x) : fderivWithin 𝕜 f (s ∩ t) x = fderivWithin 𝕜 f s x := by - have A : 𝓝[(s ∩ t) \ {x}] x = 𝓝[s \ {x}] x := by - have : (s ∩ t) \ {x} = (s \ {x}) ∩ t := by rw [inter_comm, inter_diff_assoc, inter_comm] - rw [this, ← nhdsWithin_restrict' _ ht] - simp [fderivWithin, A, hasFDerivWithinAt_inter ht] - -@[simp] -theorem fderivWithin_univ : fderivWithin 𝕜 f univ = fderiv 𝕜 f := by - ext1 x - nontriviality E - have H : 𝓝[univ \ {x}] x ≠ ⊥ := by - rw [← compl_eq_univ_diff, ← neBot_iff] - exact Module.punctured_nhds_neBot 𝕜 E x - simp [fderivWithin, fderiv, H] + simp [fderivWithin, hasFDerivWithinAt_inter ht, DifferentiableWithinAt] theorem fderivWithin_of_mem_nhds (h : s ∈ 𝓝 x) : fderivWithin 𝕜 f s x = fderiv 𝕜 f x := by rw [← fderivWithin_univ, ← univ_inter s, fderivWithin_inter h] @@ -803,11 +792,7 @@ theorem differentiableWithinAt_congr_set (h : s =ᶠ[𝓝 x] t) : theorem fderivWithin_congr_set' (y : E) (h : s =ᶠ[𝓝[{y}ᶜ] x] t) : fderivWithin 𝕜 f s x = fderivWithin 𝕜 f t x := by - have : s =ᶠ[𝓝[{x}ᶜ] x] t := nhdsWithin_compl_singleton_le x y h - have : 𝓝[s \ {x}] x = 𝓝[t \ {x}] x := by - simpa only [set_eventuallyEq_iff_inf_principal, ← nhdsWithin_inter', diff_eq, - inter_comm] using this - simp only [fderivWithin, hasFDerivWithinAt_congr_set' y h, this] + simp only [fderivWithin, differentiableWithinAt_congr_set' _ h, hasFDerivWithinAt_congr_set' _ h] theorem fderivWithin_congr_set (h : s =ᶠ[𝓝 x] t) : fderivWithin 𝕜 f s x = fderivWithin 𝕜 f t x := fderivWithin_congr_set' x <| h.filter_mono inf_le_left @@ -937,7 +922,7 @@ theorem DifferentiableWithinAt.fderivWithin_congr_mono (h : DifferentiableWithin theorem Filter.EventuallyEq.fderivWithin_eq (hs : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : fderivWithin 𝕜 f₁ s x = fderivWithin 𝕜 f s x := by - simp only [fderivWithin, hs.hasFDerivWithinAt_iff hx] + simp only [fderivWithin, DifferentiableWithinAt, hs.hasFDerivWithinAt_iff hx] theorem Filter.EventuallyEq.fderivWithin_eq_of_mem (hs : f₁ =ᶠ[𝓝[s] x] f) (hx : x ∈ s) : fderivWithin 𝕜 f₁ s x = fderivWithin 𝕜 f s x := @@ -1068,19 +1053,21 @@ theorem differentiableAt_const (c : F) : DifferentiableAt 𝕜 (fun _ => c) x := theorem differentiableWithinAt_const (c : F) : DifferentiableWithinAt 𝕜 (fun _ => c) s x := DifferentiableAt.differentiableWithinAt (differentiableAt_const _) +theorem fderivWithin_const_apply (c : F) : fderivWithin 𝕜 (fun _ => c) s x = 0 := by + rw [fderivWithin, if_pos] + apply hasFDerivWithinAt_const + +@[simp] +theorem fderivWithin_const (c : F) : fderivWithin 𝕜 (fun _ ↦ c) s = 0 := by + ext + rw [fderivWithin_const_apply, Pi.zero_apply] + theorem fderiv_const_apply (c : F) : fderiv 𝕜 (fun _ => c) x = 0 := - HasFDerivAt.fderiv (hasFDerivAt_const c x) + (hasFDerivAt_const c x).fderiv @[simp] theorem fderiv_const (c : F) : (fderiv 𝕜 fun _ : E => c) = 0 := by - ext m - rw [fderiv_const_apply] - rfl - -theorem fderivWithin_const_apply (c : F) (hxs : UniqueDiffWithinAt 𝕜 s x) : - fderivWithin 𝕜 (fun _ => c) s x = 0 := by - rw [DifferentiableAt.fderivWithin (differentiableAt_const _) hxs] - exact fderiv_const_apply _ + rw [← fderivWithin_univ, fderivWithin_const] @[simp, fun_prop] theorem differentiable_const (c : F) : Differentiable 𝕜 fun _ : E => c := fun _ => From 9e90c18d26778cd7aac98f018e6fb94c31184582 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Thu, 12 Dec 2024 18:27:08 +0000 Subject: [PATCH 683/829] feat(RingTheory/Invariant): Define invariant extensions of rings (#18560) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds a typeclass for `∀ b : B, (∀ g : G, g • b = b) → ∃ a : A, algebraMap A B a = b`, working towards Frobenius elements in #17717. --- Mathlib.lean | 1 + Mathlib/RingTheory/Invariant.lean | 92 +++++++++++++++++++++++++++++++ 2 files changed, 93 insertions(+) create mode 100644 Mathlib/RingTheory/Invariant.lean diff --git a/Mathlib.lean b/Mathlib.lean index 8d15eac376c50..4ea78cfc39c29 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4397,6 +4397,7 @@ import Mathlib.RingTheory.IntegralClosure.IsIntegral.Defs import Mathlib.RingTheory.IntegralClosure.IsIntegralClosure.Basic import Mathlib.RingTheory.IntegralClosure.IsIntegralClosure.Defs import Mathlib.RingTheory.IntegralDomain +import Mathlib.RingTheory.Invariant import Mathlib.RingTheory.IsAdjoinRoot import Mathlib.RingTheory.IsPrimary import Mathlib.RingTheory.IsTensorProduct diff --git a/Mathlib/RingTheory/Invariant.lean b/Mathlib/RingTheory/Invariant.lean new file mode 100644 index 0000000000000..46ac790026076 --- /dev/null +++ b/Mathlib/RingTheory/Invariant.lean @@ -0,0 +1,92 @@ +/- +Copyright (c) 2024 Thomas Browning. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Thomas Browning +-/ +import Mathlib.Algebra.Polynomial.GroupRingAction +import Mathlib.Algebra.Polynomial.Lifts +import Mathlib.RingTheory.IntegralClosure.Algebra.Defs + +/-! +# Invariant Extensions of Rings + +Given an extension of rings `B/A` and an action of `G` on `B`, we introduce a predicate +`Algebra.IsInvariant A B G` which states that every fixed point of `B` lies in the image of `A`. + +## Main statements +* `Algebra.IsInvariant.isIntegral`: If `G` is finite, then `B/A` is an integral extension. + +TODO: Prove the existence of Frobenius elements in this general setting (PR #17717). + +-/ + +namespace Algebra + +variable (A B G : Type*) [CommSemiring A] [Semiring B] [Algebra A B] + [Group G] [MulSemiringAction G B] + +/-- An action of a group `G` on an extension of rings `B/A` is invariant if every fixed point of +`B` lies in the image of `A`. The converse statement that every point in the image of `A` is fixed +by `G` is `smul_algebraMap` (assuming `SMulCommClass A B G`). -/ +@[mk_iff] class IsInvariant : Prop where + isInvariant : ∀ b : B, (∀ g : G, g • b = b) → ∃ a : A, algebraMap A B a = b + +end Algebra + +section IsIntegral + +variable (A B G : Type*) [CommRing A] [CommRing B] [Algebra A B] [Group G] [MulSemiringAction G B] + +namespace MulSemiringAction + +open Polynomial + +variable {B} [Fintype G] + +/-- Characteristic polynomial of a finite group action on a ring. -/ +noncomputable def charpoly (b : B) : B[X] := ∏ g : G, (X - C (g • b)) + +theorem charpoly_eq (b : B) : charpoly G b = ∏ g : G, (X - C (g • b)) := rfl + +theorem charpoly_eq_prod_smul (b : B) : charpoly G b = ∏ g : G, g • (X - C b) := by + simp only [smul_sub, smul_C, smul_X, charpoly_eq] + +theorem monic_charpoly (b : B) : (charpoly G b).Monic := + monic_prod_of_monic _ _ (fun _ _ ↦ monic_X_sub_C _) + +theorem eval_charpoly (b : B) : (charpoly G b).eval b = 0 := by + rw [charpoly_eq, eval_prod] + apply Finset.prod_eq_zero (Finset.mem_univ (1 : G)) + rw [one_smul, eval_sub, eval_C, eval_X, sub_self] + +variable {G} + +theorem smul_charpoly (b : B) (g : G) : g • (charpoly G b) = charpoly G b := by + rw [charpoly_eq_prod_smul, Finset.smul_prod_perm] + +theorem smul_coeff_charpoly (b : B) (n : ℕ) (g : G) : + g • (charpoly G b).coeff n = (charpoly G b).coeff n := by + rw [← coeff_smul, smul_charpoly] + +end MulSemiringAction + +namespace Algebra.IsInvariant + +open MulSemiringAction Polynomial + +variable [IsInvariant A B G] + +theorem charpoly_mem_lifts [Fintype G] (b : B) : + charpoly G b ∈ Polynomial.lifts (algebraMap A B) := + (charpoly G b).lifts_iff_coeff_lifts.mpr fun n ↦ isInvariant _ (smul_coeff_charpoly b n) + +theorem isIntegral [Finite G] : Algebra.IsIntegral A B := by + cases nonempty_fintype G + refine ⟨fun b ↦ ?_⟩ + obtain ⟨p, hp1, -, hp2⟩ := Polynomial.lifts_and_natDegree_eq_and_monic + (charpoly_mem_lifts A B G b) (monic_charpoly G b) + exact ⟨p, hp2, by rw [← eval_map, hp1, eval_charpoly]⟩ + +end Algebra.IsInvariant + +end IsIntegral From 3c9302f1c89690f1c68477fe9f31c228847b7c46 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 12 Dec 2024 18:27:10 +0000 Subject: [PATCH 684/829] feat(FaaDiBruno): add `norm_compAlongOrderedFinpartitionL_le` (#19713) --- .../Calculus/ContDiff/FaaDiBruno.lean | 56 +++++++++++-------- 1 file changed, 32 insertions(+), 24 deletions(-) diff --git a/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean b/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean index f8b0483d56f73..c42683bf68966 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean @@ -699,6 +699,11 @@ lemma applyOrderedFinpartition_apply (p : ∀ (i : Fin c.length), E[×c.partSize (v : Fin n → E) : c.applyOrderedFinpartition p v = (fun m ↦ p m (v ∘ c.emb m)) := rfl +theorem norm_applyOrderedFinpartition_le (p : ∀ (i : Fin c.length), E[×c.partSize i]→L[𝕜] F) + (v : Fin n → E) (m : Fin c.length) : + ‖c.applyOrderedFinpartition p v m‖ ≤ ‖p m‖ * ∏ i : Fin (c.partSize m), ‖v (c.emb m i)‖ := + (p m).le_opNorm _ + /-- Technical lemma stating how `c.applyOrderedFinpartition` commutes with updating variables. This will be the key point to show that functions constructed from `applyOrderedFinpartition` retain multilinearity. -/ @@ -741,9 +746,7 @@ variables, and for each `m` a continuous multilinear map `p m` in `c.partSize m` one can form a continuous multilinear map in `n` variables by applying `p m` to each part of the partition, and then applying `f` to the resulting vector. It is called `c.compAlongOrderedFinpartition f p`. -/ -def compAlongOrderedFinpartition - (f : ContinuousMultilinearMap 𝕜 (fun _i : Fin c.length ↦ F) G) - (p : ∀ (i : Fin c.length), E[×c.partSize i]→L[𝕜] F) : +def compAlongOrderedFinpartition (f : F [×c.length]→L[𝕜] G) (p : ∀ i, E [×c.partSize i]→L[𝕜] F) : E[×n]→L[𝕜] G where toFun v := f (c.applyOrderedFinpartition p v) map_update_add' v i x y := by @@ -757,16 +760,24 @@ def compAlongOrderedFinpartition change Continuous (fun v m ↦ p m (v ∘ c.emb m)) fun_prop -@[simp] lemma compAlongOrderFinpartition_apply - (f : ContinuousMultilinearMap 𝕜 (fun _i : Fin c.length ↦ F) G) - (p : ∀ (i : Fin c.length), E[×c.partSize i]→L[𝕜] F) (v : Fin n → E) : +@[simp] lemma compAlongOrderFinpartition_apply (f : F [×c.length]→L[𝕜] G) + (p : ∀ i, E[×c.partSize i]→L[𝕜] F) (v : Fin n → E) : c.compAlongOrderedFinpartition f p v = f (c.applyOrderedFinpartition p v) := rfl +theorem norm_compAlongOrderedFinpartition_le (f : F [×c.length]→L[𝕜] G) + (p : ∀ i, E [×c.partSize i]→L[𝕜] F) : + ‖c.compAlongOrderedFinpartition f p‖ ≤ ‖f‖ * ∏ i, ‖p i‖ := by + refine ContinuousMultilinearMap.opNorm_le_bound (by positivity) fun v ↦ ?_ + rw [compAlongOrderFinpartition_apply, mul_assoc, ← c.prod_sigma_eq_prod, + ← Finset.prod_mul_distrib] + exact f.le_opNorm_mul_prod_of_le <| c.norm_applyOrderedFinpartition_le _ _ + /-- Bundled version of `compAlongOrderedFinpartition`, depending linearly on `f` and multilinearly on `p`.-/ +@[simps apply_apply] def compAlongOrderedFinpartitionₗ : - (ContinuousMultilinearMap 𝕜 (fun _i : Fin c.length ↦ F) G) →ₗ[𝕜] - MultilinearMap 𝕜 (fun i : Fin c.length ↦ (E[×c.partSize i]→L[𝕜] F)) (E[×n]→L[𝕜] G) where + (F [×c.length]→L[𝕜] G) →ₗ[𝕜] + MultilinearMap 𝕜 (fun i : Fin c.length ↦ E[×c.partSize i]→L[𝕜] F) (E[×n]→L[𝕜] G) where toFun f := { toFun := fun p ↦ c.compAlongOrderedFinpartition f p map_update_add' := by @@ -786,24 +797,21 @@ variable (𝕜 E F G) in /-- Bundled version of `compAlongOrderedFinpartition`, depending continuously linearly on `f` and continuously multilinearly on `p`.-/ noncomputable def compAlongOrderedFinpartitionL : - (ContinuousMultilinearMap 𝕜 (fun _i : Fin c.length ↦ F) G) →L[𝕜] - ContinuousMultilinearMap 𝕜 (fun i : Fin c.length ↦ (E[×c.partSize i]→L[𝕜] F)) - (E[×n]→L[𝕜] G) := by - refine MultilinearMap.mkContinuousLinear c.compAlongOrderedFinpartitionₗ 1 (fun f p ↦ ?_) - simp only [one_mul] - change ‖c.compAlongOrderedFinpartition f p‖ ≤ _ - apply ContinuousMultilinearMap.opNorm_le_bound (by positivity) (fun v ↦ ?_) - simp only [compAlongOrderFinpartition_apply] - apply (f.le_opNorm _).trans - rw [mul_assoc, ← c.prod_sigma_eq_prod, ← Finset.prod_mul_distrib] - gcongr with m _ - exact (p m).le_opNorm _ - -@[simp] lemma compAlongOrderedFinpartitionL_apply - (f : ContinuousMultilinearMap 𝕜 (fun _i : Fin c.length ↦ F) G) + (F [×c.length]→L[𝕜] G) →L[𝕜] + ContinuousMultilinearMap 𝕜 (fun i ↦ E[×c.partSize i]→L[𝕜] F) (E[×n]→L[𝕜] G) := by + refine MultilinearMap.mkContinuousLinear c.compAlongOrderedFinpartitionₗ 1 fun f p ↦ ?_ + simp only [one_mul, compAlongOrderedFinpartitionₗ_apply_apply] + apply norm_compAlongOrderedFinpartition_le + +@[simp] lemma compAlongOrderedFinpartitionL_apply (f : F [×c.length]→L[𝕜] G) (p : ∀ (i : Fin c.length), E[×c.partSize i]→L[𝕜] F) : c.compAlongOrderedFinpartitionL 𝕜 E F G f p = c.compAlongOrderedFinpartition f p := rfl +theorem norm_compAlongOrderedFinpartitionL_le : + set_option maxSynthPendingDepth 2 in + ‖c.compAlongOrderedFinpartitionL 𝕜 E F G‖ ≤ 1 := + MultilinearMap.mkContinuousLinear_norm_le _ zero_le_one _ + end OrderedFinpartition /-! ### The Faa di Bruno formula -/ @@ -816,7 +824,7 @@ block of the composition, and then applying `q c.length` to the resulting vector called `q.compAlongComposition p c`. -/ def compAlongOrderedFinpartition {n : ℕ} (q : FormalMultilinearSeries 𝕜 F G) (p : FormalMultilinearSeries 𝕜 E F) (c : OrderedFinpartition n) : - ContinuousMultilinearMap 𝕜 (fun _i : Fin n ↦ E) G := + E [×n]→L[𝕜] G := c.compAlongOrderedFinpartition (q c.length) (fun m ↦ p (c.partSize m)) @[simp] From 2e4e286ed32fa998e9cfb2219d0a8821ab8b7155 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Thu, 12 Dec 2024 20:12:54 +0000 Subject: [PATCH 685/829] feat(SetTheory/Ordinal/Basic): recursion on well-orders (#19678) To define an ordinal-valued function, it suffices to define its values on well-ordered types. This will be used to redefine the cofinality of an ordinal as the cofinality of a well-order with its order type. --- Mathlib/SetTheory/Ordinal/Basic.lean | 40 ++++++++++++++++++++++++++++ 1 file changed, 40 insertions(+) diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index 1693d12052eee..4b873fb5d2664 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -236,16 +236,25 @@ protected theorem one_ne_zero : (1 : Ordinal) ≠ 0 := instance nontrivial : Nontrivial Ordinal.{u} := ⟨⟨1, 0, Ordinal.one_ne_zero⟩⟩ +/-- `Quotient.inductionOn` specialized to ordinals. + +Not to be confused with well-founded recursion `Ordinal.induction`. -/ @[elab_as_elim] theorem inductionOn {C : Ordinal → Prop} (o : Ordinal) (H : ∀ (α r) [IsWellOrder α r], C (type r)) : C o := Quot.inductionOn o fun ⟨α, r, wo⟩ => @H α r wo +/-- `Quotient.inductionOn₂` specialized to ordinals. + +Not to be confused with well-founded recursion `Ordinal.induction`. -/ @[elab_as_elim] theorem inductionOn₂ {C : Ordinal → Ordinal → Prop} (o₁ o₂ : Ordinal) (H : ∀ (α r) [IsWellOrder α r] (β s) [IsWellOrder β s], C (type r) (type s)) : C o₁ o₂ := Quotient.inductionOn₂ o₁ o₂ fun ⟨α, r, wo₁⟩ ⟨β, s, wo₂⟩ => @H α r wo₁ β s wo₂ +/-- `Quotient.inductionOn₃` specialized to ordinals. + +Not to be confused with well-founded recursion `Ordinal.induction`. -/ @[elab_as_elim] theorem inductionOn₃ {C : Ordinal → Ordinal → Ordinal → Prop} (o₁ o₂ o₃ : Ordinal) (H : ∀ (α r) [IsWellOrder α r] (β s) [IsWellOrder β s] (γ t) [IsWellOrder γ t], @@ -253,6 +262,37 @@ theorem inductionOn₃ {C : Ordinal → Ordinal → Ordinal → Prop} (o₁ o₂ Quotient.inductionOn₃ o₁ o₂ o₃ fun ⟨α, r, wo₁⟩ ⟨β, s, wo₂⟩ ⟨γ, t, wo₃⟩ => @H α r wo₁ β s wo₂ γ t wo₃ +open Classical in +/-- To prove a result on ordinals, it suffices to prove it for order types of well-orders. -/ +@[elab_as_elim] +theorem inductionOnWellOrder {C : Ordinal → Prop} (o : Ordinal) + (H : ∀ (α) [LinearOrder α] [WellFoundedLT α], C (typeLT α)) : C o := + inductionOn o fun α r wo ↦ @H α (linearOrderOfSTO r) wo.toIsWellFounded + +open Classical in +/-- To define a function on ordinals, it suffices to define them on order types of well-orders. + +Since `LinearOrder` is data-carrying, `liftOnWellOrder_type` is not a definitional equality, unlike +`Quotient.liftOn_mk` which is always def-eq. -/ +def liftOnWellOrder {δ : Sort v} (o : Ordinal) (f : ∀ (α) [LinearOrder α] [WellFoundedLT α], δ) + (c : ∀ (α) [LinearOrder α] [WellFoundedLT α] (β) [LinearOrder β] [WellFoundedLT β], + typeLT α = typeLT β → f α = f β) : δ := + Quotient.liftOn o (fun w ↦ @f w.α (linearOrderOfSTO w.r) w.wo.toIsWellFounded) + fun w₁ w₂ h ↦ @c + w₁.α (linearOrderOfSTO w₁.r) w₁.wo.toIsWellFounded + w₂.α (linearOrderOfSTO w₂.r) w₂.wo.toIsWellFounded + (Quotient.sound h) + +@[simp] +theorem liftOnWellOrder_type {δ : Sort v} (f : ∀ (α) [LinearOrder α] [WellFoundedLT α], δ) + (c : ∀ (α) [LinearOrder α] [WellFoundedLT α] (β) [LinearOrder β] [WellFoundedLT β], + typeLT α = typeLT β → f α = f β) {γ} [LinearOrder γ] [WellFoundedLT γ] : + liftOnWellOrder (typeLT γ) f c = f γ := by + change Quotient.liftOn' ⟦_⟧ _ _ = _ + rw [Quotient.liftOn'_mk] + congr + exact LinearOrder.ext_lt fun _ _ ↦ Iff.rfl + /-! ### The order on ordinals -/ /-- From 2cbe43aa20ec5488824a2f89d84e14e77c45d4bb Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Thu, 12 Dec 2024 20:33:25 +0000 Subject: [PATCH 686/829] feat(Algebra/Category): the category of abelian groups satisfies AB4* (#19912) --- Mathlib.lean | 2 +- Mathlib/Algebra/Category/Grp/AB.lean | 90 ++++++++++++++++++++ Mathlib/Algebra/Category/Grp/AB5.lean | 49 ----------- Mathlib/CategoryTheory/Limits/HasLimits.lean | 4 +- 4 files changed, 93 insertions(+), 52 deletions(-) create mode 100644 Mathlib/Algebra/Category/Grp/AB.lean delete mode 100644 Mathlib/Algebra/Category/Grp/AB5.lean diff --git a/Mathlib.lean b/Mathlib.lean index 4ea78cfc39c29..5845af9e4e140 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -70,7 +70,7 @@ import Mathlib.Algebra.Category.CoalgebraCat.ComonEquivalence import Mathlib.Algebra.Category.CoalgebraCat.Monoidal import Mathlib.Algebra.Category.FGModuleCat.Basic import Mathlib.Algebra.Category.FGModuleCat.Limits -import Mathlib.Algebra.Category.Grp.AB5 +import Mathlib.Algebra.Category.Grp.AB import Mathlib.Algebra.Category.Grp.Abelian import Mathlib.Algebra.Category.Grp.Adjunctions import Mathlib.Algebra.Category.Grp.Basic diff --git a/Mathlib/Algebra/Category/Grp/AB.lean b/Mathlib/Algebra/Category/Grp/AB.lean new file mode 100644 index 0000000000000..edbbcb51bdb57 --- /dev/null +++ b/Mathlib/Algebra/Category/Grp/AB.lean @@ -0,0 +1,90 @@ +/- +Copyright (c) 2023 David Kurniadi Angdinata. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Kurniadi Angdinata, Moritz Firsching, Nikolas Kuhn, Amelia Livingston +-/ +import Mathlib.Algebra.Category.Grp.Biproducts +import Mathlib.Algebra.Category.Grp.FilteredColimits +import Mathlib.Algebra.Homology.ShortComplex.Ab +import Mathlib.CategoryTheory.Abelian.GrothendieckAxioms +import Mathlib.CategoryTheory.Limits.FunctorCategory.EpiMono +/-! +# AB axioms for the category of abelian groups + +This file proves that the category of abelian groups satisfies Grothendieck's axioms AB5, AB4, and +AB4*. +-/ + +universe u + +open CategoryTheory Limits + +instance {J C : Type*} [Category J] [Category C] [HasColimitsOfShape J C] [Preadditive C] : + (colim (J := J) (C := C)).Additive where + +variable {J : Type u} [SmallCategory J] [IsFiltered J] + +noncomputable instance : + (colim (J := J) (C := AddCommGrp.{u})).PreservesHomology := + Functor.preservesHomology_of_map_exact _ (fun S hS ↦ by + replace hS := fun j => hS.map ((evaluation _ _).obj j) + simp only [ShortComplex.ab_exact_iff_ker_le_range] at hS ⊢ + intro x (hx : _ = _) + dsimp at hx + rcases Concrete.colimit_exists_rep S.X₂ x with ⟨j, y, rfl⟩ + erw [← comp_apply, colimit.ι_map, comp_apply, + ← map_zero (by exact colimit.ι S.X₃ j : (S.X₃).obj j →+ ↑(colimit S.X₃))] at hx + rcases Concrete.colimit_exists_of_rep_eq.{u, u, u} S.X₃ _ _ hx + with ⟨k, e₁, e₂, hk : _ = S.X₃.map e₂ 0⟩ + rw [map_zero, ← comp_apply, ← NatTrans.naturality, comp_apply] at hk + rcases hS k hk with ⟨t, ht⟩ + use colimit.ι S.X₁ k t + erw [← comp_apply, colimit.ι_map, comp_apply, ht] + exact colimit.w_apply S.X₂ e₁ y) + +noncomputable instance : + PreservesFiniteLimits <| colim (J := J) (C := AddCommGrp.{u}) := by + apply Functor.preservesFiniteLimits_of_preservesHomology + +instance : HasFilteredColimits (AddCommGrp.{u}) where + HasColimitsOfShape := inferInstance + +noncomputable instance : AB5 (AddCommGrp.{u}) where + ofShape _ := { preservesFiniteLimits := inferInstance } + +attribute [local instance] Abelian.hasFiniteBiproducts + +instance : AB4 AddCommGrp.{u} := AB4.of_AB5 _ + +instance : HasExactLimitsOfShape (Discrete J) (AddCommGrp.{u}) := by + apply ( config := {allowSynthFailures := true} ) hasExactLimitsOfShape_of_preservesEpi + exact { + preserves {X Y} f hf := by + let iX : limit X ≅ AddCommGrp.of ((i : J) → X.obj ⟨i⟩) := (Pi.isoLimit X).symm ≪≫ + (limit.isLimit _).conePointUniqueUpToIso (AddCommGrp.HasLimit.productLimitCone _).isLimit + let iY : limit Y ≅ AddCommGrp.of ((i : J) → Y.obj ⟨i⟩) := (Pi.isoLimit Y).symm ≪≫ + (limit.isLimit _).conePointUniqueUpToIso (AddCommGrp.HasLimit.productLimitCone _).isLimit + have : Pi.map (fun i ↦ f.app ⟨i⟩) = iX.inv ≫ lim.map f ≫ iY.hom := by + simp only [AddCommGrp.coe_of, Functor.comp_obj, Discrete.functor_obj_eq_as, Discrete.mk_as, + Pi.isoLimit, IsLimit.conePointUniqueUpToIso, limit.cone, + AddCommGrp.HasLimit.productLimitCone, Iso.trans_inv, Functor.mapIso_inv, + IsLimit.uniqueUpToIso_inv, Cones.forget_map, IsLimit.liftConeMorphism_hom, + limit.isLimit_lift, Iso.symm_inv, Functor.mapIso_hom, IsLimit.uniqueUpToIso_hom, lim_obj, + lim_map, Iso.trans_hom, Iso.symm_hom, AddCommGrp.HasLimit.lift, Functor.const_obj_obj, + Category.assoc, limit.lift_map_assoc, Pi.cone_pt, iX, iY] + ext g j + change _ = (_ ≫ limit.π (Discrete.functor fun j ↦ Y.obj { as := j }) ⟨j⟩) _ + simp only [Discrete.functor_obj_eq_as, Functor.comp_obj, Discrete.mk_as, productIsProduct', + limit.lift_π, Fan.mk_pt, Fan.mk_π_app, Pi.map_apply] + change _ = (_ ≫ _ ≫ limit.π Y ⟨j⟩) _ + simp + suffices Epi (iX.hom ≫ (iX.inv ≫ lim.map f ≫ iY.hom) ≫ iY.inv) by simpa using this + suffices Epi (iX.inv ≫ lim.map f ≫ iY.hom) from inferInstance + rw [AddCommGrp.epi_iff_surjective, ← this] + simp_rw [CategoryTheory.NatTrans.epi_iff_epi_app, AddCommGrp.epi_iff_surjective] at hf + refine fun b ↦ ⟨fun i ↦ (hf ⟨i⟩ (b i)).choose, ?_⟩ + funext i + exact (hf ⟨i⟩ (b i)).choose_spec } + +instance : AB4Star AddCommGrp.{u} where + ofShape _ := inferInstance diff --git a/Mathlib/Algebra/Category/Grp/AB5.lean b/Mathlib/Algebra/Category/Grp/AB5.lean deleted file mode 100644 index 81a72475dbf30..0000000000000 --- a/Mathlib/Algebra/Category/Grp/AB5.lean +++ /dev/null @@ -1,49 +0,0 @@ -/- -Copyright (c) 2023 David Kurniadi Angdinata. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: David Kurniadi Angdinata, Moritz Firsching, Nikolas Kuhn, Amelia Livingston --/ -import Mathlib.Algebra.Category.Grp.FilteredColimits -import Mathlib.Algebra.Homology.ShortComplex.Ab -import Mathlib.CategoryTheory.Abelian.GrothendieckAxioms -/-! -# The category of abelian groups satisfies Grothendieck's axiom AB5 - --/ - -universe u - -open CategoryTheory Limits - -instance {J C : Type*} [Category J] [Category C] [HasColimitsOfShape J C] [Preadditive C] : - (colim (J := J) (C := C)).Additive where - -variable {J : Type u} [SmallCategory J] [IsFiltered J] - -noncomputable instance : - (colim (J := J) (C := AddCommGrp.{u})).PreservesHomology := - Functor.preservesHomology_of_map_exact _ (fun S hS ↦ by - replace hS := fun j => hS.map ((evaluation _ _).obj j) - simp only [ShortComplex.ab_exact_iff_ker_le_range] at hS ⊢ - intro x (hx : _ = _) - dsimp at hx - rcases Concrete.colimit_exists_rep S.X₂ x with ⟨j, y, rfl⟩ - erw [← comp_apply, colimit.ι_map, comp_apply, - ← map_zero (by exact colimit.ι S.X₃ j : (S.X₃).obj j →+ ↑(colimit S.X₃))] at hx - rcases Concrete.colimit_exists_of_rep_eq.{u, u, u} S.X₃ _ _ hx - with ⟨k, e₁, e₂, hk : _ = S.X₃.map e₂ 0⟩ - rw [map_zero, ← comp_apply, ← NatTrans.naturality, comp_apply] at hk - rcases hS k hk with ⟨t, ht⟩ - use colimit.ι S.X₁ k t - erw [← comp_apply, colimit.ι_map, comp_apply, ht] - exact colimit.w_apply S.X₂ e₁ y) - -noncomputable instance : - PreservesFiniteLimits <| colim (J := J) (C := AddCommGrp.{u}) := by - apply Functor.preservesFiniteLimits_of_preservesHomology - -instance : HasFilteredColimits (AddCommGrp.{u}) where - HasColimitsOfShape := inferInstance - -noncomputable instance : AB5 (AddCommGrp.{u}) where - ofShape _ := { preservesFiniteLimits := inferInstance } diff --git a/Mathlib/CategoryTheory/Limits/HasLimits.lean b/Mathlib/CategoryTheory/Limits/HasLimits.lean index 744fd90f39333..27bcb80063f88 100644 --- a/Mathlib/CategoryTheory/Limits/HasLimits.lean +++ b/Mathlib/CategoryTheory/Limits/HasLimits.lean @@ -238,7 +238,7 @@ theorem limit.hom_ext {F : J ⥤ C} [HasLimit F] {X : C} {f f' : X ⟶ limit F} (w : ∀ j, f ≫ limit.π F j = f' ≫ limit.π F j) : f = f' := (limit.isLimit F).hom_ext w -@[simp] +@[reassoc (attr := simp)] theorem limit.lift_map {F G : J ⥤ C} [HasLimit F] [HasLimit G] (c : Cone F) (α : F ⟶ G) : limit.lift F c ≫ limMap α = limit.lift G ((Cones.postcompose α).obj c) := by ext @@ -995,7 +995,7 @@ variable {G : J ⥤ C} (α : F ⟶ G) @[reassoc] theorem colimit.ι_map (j : J) : colimit.ι F j ≫ colim.map α = α.app j ≫ colimit.ι G j := by simp -@[simp] +@[reassoc (attr := simp)] theorem colimit.map_desc (c : Cocone G) : colimMap α ≫ colimit.desc G c = colimit.desc F ((Cocones.precompose α).obj c) := by ext j From 815cff67b7eef7788f91a6f49ff6dfadbe1cd933 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Thu, 12 Dec 2024 20:33:26 +0000 Subject: [PATCH 687/829] chore: simp `isEmpty_coe_sort` instead of `isEmpty_subtype` (#19918) See [Zulip](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/isEmpty_subtype/near/485723811). --- Mathlib/Data/Set/Basic.lean | 3 +-- Mathlib/Logic/IsEmpty.lean | 1 - Mathlib/Order/CompleteBooleanAlgebra.lean | 4 ++-- Mathlib/SetTheory/Cardinal/Finite.lean | 2 +- Mathlib/Topology/Instances/AddCircle.lean | 2 +- 5 files changed, 5 insertions(+), 7 deletions(-) diff --git a/Mathlib/Data/Set/Basic.lean b/Mathlib/Data/Set/Basic.lean index ae281b9e0c8a5..7a6434c8186c5 100644 --- a/Mathlib/Data/Set/Basic.lean +++ b/Mathlib/Data/Set/Basic.lean @@ -502,8 +502,7 @@ alias ⟨Nonempty.ne_empty, _⟩ := nonempty_iff_ne_empty @[simp] theorem not_nonempty_empty : ¬(∅ : Set α).Nonempty := fun ⟨_, hx⟩ => hx --- Porting note: removing `@[simp]` as it is competing with `isEmpty_subtype`. --- @[simp] +@[simp] theorem isEmpty_coe_sort {s : Set α} : IsEmpty (↥s) ↔ s = ∅ := not_iff_not.1 <| by simpa using nonempty_iff_ne_empty diff --git a/Mathlib/Logic/IsEmpty.lean b/Mathlib/Logic/IsEmpty.lean index aad5c6c29d0ee..40fbad5e8c46a 100644 --- a/Mathlib/Logic/IsEmpty.lean +++ b/Mathlib/Logic/IsEmpty.lean @@ -160,7 +160,6 @@ theorem isEmpty_sigma {α} {E : α → Type*} : IsEmpty (Sigma E) ↔ ∀ a, IsE theorem isEmpty_psigma {α} {E : α → Sort*} : IsEmpty (PSigma E) ↔ ∀ a, IsEmpty (E a) := by simp only [← not_nonempty_iff, nonempty_psigma, not_exists] -@[simp] theorem isEmpty_subtype (p : α → Prop) : IsEmpty (Subtype p) ↔ ∀ x, ¬p x := by simp only [← not_nonempty_iff, nonempty_subtype, not_exists] diff --git a/Mathlib/Order/CompleteBooleanAlgebra.lean b/Mathlib/Order/CompleteBooleanAlgebra.lean index 352874a908984..bc8dde3dbae38 100644 --- a/Mathlib/Order/CompleteBooleanAlgebra.lean +++ b/Mathlib/Order/CompleteBooleanAlgebra.lean @@ -269,7 +269,7 @@ abbrev toCompleteDistribLattice : CompleteDistribLattice.MinimalAxioms α where _ = ⨅ i : ULift.{u} Bool, ⨆ j : match i with | .up true => PUnit.{u + 1} | .up false => s, match i with | .up true => a - | .up false => j := by simp [le_antisymm_iff, sSup_eq_iSup', iSup_unique, iInf_bool_eq] + | .up false => j := by simp [sSup_eq_iSup', iSup_unique, iInf_bool_eq] _ ≤ _ := by simp only [minAx.iInf_iSup_eq, iInf_ulift, iInf_bool_eq, iSup_le_iff] exact fun x ↦ le_biSup _ (x (.up false)).2 @@ -282,7 +282,7 @@ abbrev toCompleteDistribLattice : CompleteDistribLattice.MinimalAxioms α where | .up false => j := by simp only [minAx.iSup_iInf_eq, iSup_ulift, iSup_bool_eq, le_iInf_iff] exact fun x ↦ biInf_le _ (x (.up false)).2 - _ = _ := by simp [le_antisymm_iff, sInf_eq_iInf', iInf_unique, iSup_bool_eq] + _ = _ := by simp [sInf_eq_iInf', iInf_unique, iSup_bool_eq] /-- The `CompletelyDistribLattice.MinimalAxioms` element corresponding to a frame. -/ def of [CompletelyDistribLattice α] : MinimalAxioms α := { ‹CompletelyDistribLattice α› with } diff --git a/Mathlib/SetTheory/Cardinal/Finite.lean b/Mathlib/SetTheory/Cardinal/Finite.lean index ed8bad925692c..5de36466bcce0 100644 --- a/Mathlib/SetTheory/Cardinal/Finite.lean +++ b/Mathlib/SetTheory/Cardinal/Finite.lean @@ -236,7 +236,7 @@ lemma card_prod_singleton (s : Set α) (b : β) : Nat.card (s ×ˢ {b}) = Nat.ca rw [prod_singleton, Nat.card_image_of_injective (Prod.mk.inj_right b)] theorem natCard_pos (hs : s.Finite) : 0 < Nat.card s ↔ s.Nonempty := by - simp [pos_iff_ne_zero, Nat.card_eq_zero, hs.to_subtype, Set.nonempty_def, nonempty_iff_ne_empty] + simp [pos_iff_ne_zero, Nat.card_eq_zero, hs.to_subtype, nonempty_iff_ne_empty] protected alias ⟨_, Nonempty.natCard_pos⟩ := natCard_pos diff --git a/Mathlib/Topology/Instances/AddCircle.lean b/Mathlib/Topology/Instances/AddCircle.lean index 3a4b8e7bacfd0..19d581ed88212 100644 --- a/Mathlib/Topology/Instances/AddCircle.lean +++ b/Mathlib/Topology/Instances/AddCircle.lean @@ -478,7 +478,7 @@ theorem card_addOrderOf_eq_totient {n : ℕ} : erw [infinite_coe_iff] exact infinite_not_isOfFinAddOrder hu exact Nat.card_eq_zero_of_infinite - · have : IsEmpty { u : AddCircle p // ¬IsOfFinAddOrder u } := by simpa using h + · have : IsEmpty { u : AddCircle p // ¬IsOfFinAddOrder u } := by simpa [isEmpty_subtype] using h exact Nat.card_of_isEmpty · rw [← coe_setOf, Nat.card_congr (setAddOrderOfEquiv p hn), n.totient_eq_card_lt_and_coprime] From d554797db161743e93778a10b0fe9336dce8527a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Thu, 12 Dec 2024 22:51:27 +0000 Subject: [PATCH 688/829] chore(SetTheory/Cardinal/Arithmetic): generalize `card_iSup_Iio_le_sum_card` (#19727) The theorem becomes strictly more general at zero cost by letting the domain of the function be `Iio o` rather than `Ordinal`. --- Mathlib/SetTheory/Cardinal/Arithmetic.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/SetTheory/Cardinal/Arithmetic.lean b/Mathlib/SetTheory/Cardinal/Arithmetic.lean index 8e9e3236e07e2..8f3b807c4c581 100644 --- a/Mathlib/SetTheory/Cardinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Cardinal/Arithmetic.lean @@ -869,16 +869,16 @@ theorem card_iSup_le_sum_card {ι : Type u} (f : ι → Ordinal.{max u v}) : have := lift_card_iSup_le_sum_card f rwa [Cardinal.lift_id'] at this -theorem card_iSup_Iio_le_sum_card {o : Ordinal.{u}} (f : Ordinal.{u} → Ordinal.{max u v}) : +theorem card_iSup_Iio_le_sum_card {o : Ordinal.{u}} (f : Iio o → Ordinal.{max u v}) : (⨆ a : Iio o, f a).card ≤ Cardinal.sum fun i ↦ (f ((enumIsoToType o).symm i)).card := by apply le_of_eq_of_le (congr_arg _ _).symm (card_iSup_le_sum_card _) - simpa using (enumIsoToType o).symm.iSup_comp (g := fun x ↦ f x.1) + simpa using (enumIsoToType o).symm.iSup_comp (g := fun x ↦ f x) -theorem card_iSup_Iio_le_card_mul_iSup {o : Ordinal.{u}} (f : Ordinal.{u} → Ordinal.{max u v}) : +theorem card_iSup_Iio_le_card_mul_iSup {o : Ordinal.{u}} (f : Iio o → Ordinal.{max u v}) : (⨆ a : Iio o, f a).card ≤ Cardinal.lift.{v} o.card * ⨆ a : Iio o, (f a).card := by apply (card_iSup_Iio_le_sum_card f).trans convert ← sum_le_iSup_lift _ · exact mk_toType o - · exact (enumIsoToType o).symm.iSup_comp (g := fun x ↦ (f x.1).card) + · exact (enumIsoToType o).symm.iSup_comp (g := fun x ↦ (f x).card) end Ordinal From efc0961e8c56d38dceb9cbacc1330aeba8a41d9a Mon Sep 17 00:00:00 2001 From: Rob Lewis Date: Fri, 13 Dec 2024 00:19:49 +0000 Subject: [PATCH 689/829] chore: add badge to open mathlib in Codespaces (#19927) --- README.md | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/README.md b/README.md index 5dd59489b323e..ab1e1d8edd48e 100644 --- a/README.md +++ b/README.md @@ -12,7 +12,9 @@ as well as tactics that use the former and allow to develop the latter. ## Installation You can find detailed instructions to install Lean, mathlib, and supporting tools on [our website](https://leanprover-community.github.io/get_started.html). -Alternatively, click on the button below to open a Gitpod workspace containing the project. +Alternatively, click on one of the buttons below to open a GitHub Codespace or a Gitpod workspace containing the project. + +[![Open in GitHub Codespaces](https://github.com/codespaces/badge.svg)](https://codespaces.new/leanprover-community/mathlib4) [![Open in Gitpod](https://gitpod.io/button/open-in-gitpod.svg)](https://gitpod.io/#https://github.com/leanprover-community/mathlib4) From a1c80a59b69da977f8519230ef96b3b95e31bc1d Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 13 Dec 2024 00:33:05 +0000 Subject: [PATCH 690/829] chore: further backports from leanprover/lean4#6123 (#19929) Follow-up from #19852. https://github.com/leanprover/lean4/pull/6123 fixes some inconsistencies in the behaviour of simp, by propagating Simp.Config settings through to Meta.Config. This requires us to be more explicit about unfolding let bindings in many places in Mathlib. There are a few more changes that will only land in `nightly-testing` once https://github.com/leanprover/lean4/pull/6123 is in. --- .../AscendingDescendingSequences.lean | 10 ++++--- Archive/Wiedijk100Theorems/CubingACube.lean | 20 ++++++------- Archive/Wiedijk100Theorems/HeronsFormula.lean | 5 ++-- .../SolutionOfCubicQuartic.lean | 6 ++-- Archive/ZagierTwoSquares.lean | 2 +- Counterexamples/DirectSumIsInternal.lean | 2 +- Counterexamples/Phillips.lean | 6 ++-- Mathlib/Algebra/CharZero/Quotient.lean | 2 +- .../Computation/Approximations.lean | 4 +-- .../Computation/CorrectnessTerminating.lean | 2 +- .../Computation/TerminatesIffRat.lean | 4 +-- .../ContinuedFractions/ConvergentsEquiv.lean | 2 +- Mathlib/Algebra/GroupWithZero/WithZero.lean | 5 +++- .../Algebra/Homology/ShortComplex/Exact.lean | 4 +-- Mathlib/Algebra/Lie/TraceForm.lean | 3 +- Mathlib/Algebra/Lie/Weights/Basic.lean | 2 +- Mathlib/Algebra/Module/FreeLocus.lean | 2 +- .../Morphisms/RingHomProperties.lean | 2 +- Mathlib/Analysis/Convex/Jensen.lean | 2 +- .../FunctionalSpaces/SobolevInequality.lean | 8 ++--- Mathlib/Analysis/InnerProductSpace/PiL2.lean | 2 +- .../Analysis/InnerProductSpace/TwoDim.lean | 8 ++--- Mathlib/CategoryTheory/Functor/Flat.lean | 2 ++ .../Sites/DenseSubsite/SheafEquiv.lean | 2 +- Mathlib/Combinatorics/Colex.lean | 3 +- Mathlib/Combinatorics/Hall/Basic.lean | 7 ++--- .../SimpleGraph/Regularity/Equitabilise.lean | 2 +- Mathlib/Computability/PartrecCode.lean | 12 ++++---- Mathlib/Computability/Primrec.lean | 2 +- Mathlib/Data/ENat/Basic.lean | 4 ++- Mathlib/Data/Matrix/Rank.lean | 4 +-- Mathlib/Data/Nat/Fib/Zeckendorf.lean | 2 +- Mathlib/Data/Nat/ModEq.lean | 4 +-- Mathlib/Data/Nat/Pairing.lean | 6 ++-- Mathlib/Data/Nat/Prime/Basic.lean | 4 +-- Mathlib/Data/Nat/Prime/Defs.lean | 7 +++-- Mathlib/Data/Nat/Sqrt.lean | 4 +-- Mathlib/FieldTheory/PurelyInseparable.lean | 2 +- .../Geometry/Manifold/ContMDiff/Atlas.lean | 30 ++++++++++--------- .../Geometry/Manifold/ContMDiff/Basic.lean | 9 +++--- .../LinearAlgebra/Matrix/Charpoly/Coeff.lean | 2 +- Mathlib/LinearAlgebra/SesquilinearForm.lean | 2 +- Mathlib/MeasureTheory/Constructions/Pi.lean | 2 +- .../Covering/Differentiation.lean | 8 ++--- .../Integral/VitaliCaratheodory.lean | 2 +- .../Measure/Lebesgue/VolumeOfBalls.lean | 4 +-- Mathlib/NumberTheory/Bernoulli.lean | 2 +- Mathlib/NumberTheory/EulerProduct/Basic.lean | 2 +- Mathlib/NumberTheory/WellApproximable.lean | 4 +-- Mathlib/Order/Filter/Basic.lean | 2 +- Mathlib/Order/LiminfLimsup.lean | 6 ++-- Mathlib/Order/Partition/Equipartition.lean | 2 +- Mathlib/Probability/Integration.lean | 1 + Mathlib/Probability/Process/Stopping.lean | 2 +- Mathlib/Probability/StrongLaw.lean | 4 +-- Mathlib/RingTheory/FinitePresentation.lean | 2 +- Mathlib/RingTheory/Flat/Basic.lean | 4 +-- Mathlib/RingTheory/Jacobson/Ring.lean | 6 ++-- Mathlib/RingTheory/Localization/Ideal.lean | 2 +- .../Symmetric/NewtonIdentities.lean | 4 +-- .../Polynomial/Eisenstein/IsIntegral.lean | 2 +- Mathlib/RingTheory/Polynomial/Nilpotent.lean | 6 ++-- Mathlib/RingTheory/RingHom/Locally.lean | 6 ++-- Mathlib/RingTheory/Smooth/Kaehler.lean | 2 +- Mathlib/RingTheory/Trace/Quotient.lean | 4 +-- Mathlib/RingTheory/Unramified/Field.lean | 6 ++-- Mathlib/RingTheory/Unramified/Finite.lean | 3 +- .../Topology/Algebra/InfiniteSum/Real.lean | 4 +-- Mathlib/Topology/Category/Stonean/Basic.lean | 2 +- .../Topology/MetricSpace/GromovHausdorff.lean | 16 +++++----- Mathlib/Topology/MetricSpace/Polish.lean | 2 +- .../UniformSpace/CompactConvergence.lean | 4 +-- 72 files changed, 173 insertions(+), 156 deletions(-) diff --git a/Archive/Wiedijk100Theorems/AscendingDescendingSequences.lean b/Archive/Wiedijk100Theorems/AscendingDescendingSequences.lean index 884fb11207295..8f23a38188c87 100644 --- a/Archive/Wiedijk100Theorems/AscendingDescendingSequences.lean +++ b/Archive/Wiedijk100Theorems/AscendingDescendingSequences.lean @@ -72,10 +72,10 @@ theorem erdos_szekeres {r s n : ℕ} {f : Fin n → α} (hn : r * s < n) (hf : I · refine Or.imp ?_ ?_ hi on_goal 1 => have : (ab i).1 ∈ image card (inc_sequences_ending_in i) := by - simp only [← hab]; exact max'_mem _ _ + simp only [ab', ← hab]; exact max'_mem _ _ on_goal 2 => have : (ab i).2 ∈ image card (dec_sequences_ending_in i) := by - simp only [← hab]; exact max'_mem _ _ + simp only [ab', ← hab]; exact max'_mem _ _ all_goals intro hi rw [mem_image] at this @@ -94,10 +94,12 @@ theorem erdos_szekeres {r s n : ℕ} {f : Fin n → α} (hn : r * s < n) (hf : I cases lt_or_gt_of_ne fun _ => ne_of_lt ‹i < j› (hf ‹f i = f j›) on_goal 1 => apply ne_of_lt _ q₁ - have : (ab' i).1 ∈ image card (inc_sequences_ending_in i) := by dsimp only; exact max'_mem _ _ + have : (ab' i).1 ∈ image card (inc_sequences_ending_in i) := by + dsimp only [ab']; exact max'_mem _ _ on_goal 2 => apply ne_of_lt _ q₂ - have : (ab' i).2 ∈ image card (dec_sequences_ending_in i) := by dsimp only; exact max'_mem _ _ + have : (ab' i).2 ∈ image card (dec_sequences_ending_in i) := by + dsimp only [ab']; exact max'_mem _ _ all_goals -- Reduce to showing there is a subsequence of length `a_i + 1` which ends at `j`. rw [Nat.lt_iff_add_one_le] diff --git a/Archive/Wiedijk100Theorems/CubingACube.lean b/Archive/Wiedijk100Theorems/CubingACube.lean index 6a9ccd178470c..2c877b2adc27d 100644 --- a/Archive/Wiedijk100Theorems/CubingACube.lean +++ b/Archive/Wiedijk100Theorems/CubingACube.lean @@ -277,7 +277,7 @@ theorem nontrivial_bcubes : (bcubes cs c).Nontrivial := by have h2i : i ∈ bcubes cs c := ⟨hi.1.symm, v.2.1 i hi.1.symm ⟨tail c.b, hi.2, fun j => c.b_mem_side j.succ⟩⟩ let j : Fin (n + 1) := ⟨2, h.three_le⟩ - have hj : 0 ≠ j := by simp only [Fin.ext_iff, Ne]; norm_num + have hj : 0 ≠ j := by simp only [j, Fin.ext_iff, Ne]; norm_num let p : Fin (n + 1) → ℝ := fun j' => if j' = j then c.b j + (cs i).w else c.b j' have hp : p ∈ c.bottom := by constructor @@ -343,13 +343,13 @@ theorem smallest_onBoundary {j} (bi : OnBoundary (mi_mem_bcubes : mi h v ∈ _) let i := mi h v; have hi : i ∈ bcubes cs c := mi_mem_bcubes cases' bi with bi bi · refine ⟨(cs i).b j.succ + (cs i).w, ⟨?_, ?_⟩, ?_⟩ - · simp [side, bi, hw', w_lt_w h v hi] - · intro h'; simpa [lt_irrefl] using h'.2 + · simp [i, side, bi, hw', w_lt_w h v hi] + · intro h'; simpa [i, lt_irrefl] using h'.2 intro i' hi' i'_i h2i'; constructor · apply le_trans h2i'.1 - simp [hw'] + simp [i, hw'] apply lt_of_lt_of_le (add_lt_add_left (mi_strict_minimal i'_i.symm hi') _) - simp [bi.symm, b_le_b hi'] + simp [i, bi.symm, b_le_b hi'] let s := bcubes cs c \ {i} have hs : s.Nonempty := by rcases (nontrivial_bcubes h v).exists_ne i with ⟨i', hi', h2i'⟩ @@ -368,7 +368,7 @@ theorem smallest_onBoundary {j} (bi : OnBoundary (mi_mem_bcubes : mi h v ∈ _) intro i'' hi'' h2i'' h3i''; constructor; swap; · apply lt_trans hx h3i''.2 rw [le_sub_iff_add_le] refine le_trans ?_ (t_le_t hi'' j); rw [add_le_add_iff_left]; apply h3i' i'' ⟨hi'', _⟩ - simp [mem_singleton, h2i''] + simp [i, mem_singleton, h2i''] variable (h v) @@ -390,7 +390,7 @@ theorem mi_not_onBoundary (j : Fin n) : ¬OnBoundary (mi_mem_bcubes : mi h v ∈ simp only [hj₂, if_false]; apply tail_sub hi; apply b_mem_side rcases v.1 hp with ⟨_, ⟨i', rfl⟩, hi'⟩ have h2i' : i' ∈ bcubes cs c := ⟨hi'.1.symm, v.2.1 i' hi'.1.symm ⟨tail p, hi'.2, hp.2⟩⟩ - have i_i' : i ≠ i' := by rintro rfl; simpa [p, side_tail, h2x] using hi'.2 j + have i_i' : i ≠ i' := by rintro rfl; simpa [i, p, side_tail, h2x] using hi'.2 j have : Nonempty (↥((cs i').tail.side j' \ (cs i).tail.side j')) := by apply nonempty_Ico_sdiff · apply mi_strict_minimal i_i' h2i' @@ -420,8 +420,8 @@ theorem mi_not_onBoundary (j : Fin n) : ¬OnBoundary (mi_mem_bcubes : mi h v ∈ intro j₂ by_cases hj₂ : j₂ = j · cases hj₂; refine ⟨x, ?_, ?_⟩ - · convert hi'.2 j using 1; simp [p] - apply h3x h2i'' i_i''.symm; convert hi''.2 j using 1; simp [p', hj'.symm] + · convert hi'.2 j using 1; simp [i, p] + apply h3x h2i'' i_i''.symm; convert hi''.2 j using 1; simp [i, p', hj'.symm] by_cases h2j₂ : j₂ = j' · cases h2j₂; refine ⟨x', hx'.1, ?_⟩; convert hi''.2 j' using 1; simp [p'] refine ⟨(cs i).b j₂.succ, ?_, ?_⟩ @@ -480,7 +480,7 @@ theorem valley_mi : Valley cs (cs (mi h v)).shiftUp := by have h3i'' : (cs i).w < (cs i'').w := by apply mi_strict_minimal _ h2i''; rintro rfl; apply h2p3; convert hi''.2 let p' := @cons n (fun _ => ℝ) (cs i).xm p3 - have hp' : p' ∈ (cs i').toSet := by simpa [p', toSet, forall_iff_succ, hi'.symm] using h1p3 + have hp' : p' ∈ (cs i').toSet := by simpa [i, p', toSet, forall_iff_succ, hi'.symm] using h1p3 have h2p' : p' ∈ (cs i'').toSet := by simp only [p', toSet, forall_iff_succ, cons_succ, cons_zero, mem_setOf_eq] refine ⟨?_, by simpa [toSet] using hi''.2⟩ diff --git a/Archive/Wiedijk100Theorems/HeronsFormula.lean b/Archive/Wiedijk100Theorems/HeronsFormula.lean index e64842f51e9ae..441fd3bcdf092 100644 --- a/Archive/Wiedijk100Theorems/HeronsFormula.lean +++ b/Archive/Wiedijk100Theorems/HeronsFormula.lean @@ -45,11 +45,12 @@ theorem heron {p1 p2 p3 : P} (h1 : p1 ≠ p2) (h2 : p3 ≠ p2) : let γ := ∠ p1 p2 p3 obtain := (dist_pos.mpr h1).ne', (dist_pos.mpr h2).ne' have cos_rule : cos γ = (a * a + b * b - c * c) / (2 * a * b) := by - field_simp [mul_comm, + field_simp [a, b, c, γ, mul_comm, dist_sq_eq_dist_sq_add_dist_sq_sub_two_mul_dist_mul_dist_mul_cos_angle p1 p2 p3] let numerator := (2 * a * b) ^ 2 - (a * a + b * b - c * c) ^ 2 let denominator := (2 * a * b) ^ 2 - have split_to_frac : ↑1 - cos γ ^ 2 = numerator / denominator := by field_simp [cos_rule] + have split_to_frac : ↑1 - cos γ ^ 2 = numerator / denominator := by + field_simp [numerator, denominator, cos_rule] have numerator_nonneg : 0 ≤ numerator := by have frac_nonneg : 0 ≤ numerator / denominator := (sub_nonneg.mpr (cos_sq_le_one γ)).trans_eq split_to_frac diff --git a/Archive/Wiedijk100Theorems/SolutionOfCubicQuartic.lean b/Archive/Wiedijk100Theorems/SolutionOfCubicQuartic.lean index 32528a91ed9fe..5b2f6590644dd 100644 --- a/Archive/Wiedijk100Theorems/SolutionOfCubicQuartic.lean +++ b/Archive/Wiedijk100Theorems/SolutionOfCubicQuartic.lean @@ -119,7 +119,7 @@ theorem cubic_eq_zero_iff (ha : a ≠ 0) (hω : IsPrimitiveRoot ω 3) simp [field_simps, y, ha, h9, h54]; ring have h₂ : ∀ x, a * x = 0 ↔ x = 0 := by intro x; simp [ha] rw [h₁, h₂, cubic_depressed_eq_zero_iff hω hp_nonzero hr hs3 ht] - simp_rw [eq_sub_iff_add_eq] + simp_rw [y, eq_sub_iff_add_eq] /-- The solution of the cubic equation when p equals zero. -/ theorem cubic_eq_zero_iff_of_p_eq_zero (ha : a ≠ 0) (hω : IsPrimitiveRoot ω 3) @@ -218,7 +218,7 @@ theorem quartic_eq_zero_iff (ha : a ≠ 0) simp [field_simps, y, ha, h4, h8, h16, h256]; ring have h₂ : ∀ x, a * x = 0 ↔ x = 0 := by intro x; simp [ha] rw [h₁, h₂, quartic_depressed_eq_zero_iff hq_nonzero hu hs hv hw] - simp_rw [eq_sub_iff_add_eq] + simp_rw [y, eq_sub_iff_add_eq] /-- The roots of a quartic polynomial when `q` equals zero. -/ theorem quartic_eq_zero_iff_of_q_eq_zero (ha : a ≠ 0) @@ -249,7 +249,7 @@ theorem quartic_eq_zero_iff_of_q_eq_zero (ha : a ≠ 0) have ht' : discrim 1 p r = t * t := by rw [discrim]; linear_combination -ht rw [quadratic_eq_zero_iff one_ne_zero ht', mul_one] _ ↔ _ := by - simp_rw [← hv, ← hw, pow_two, mul_self_eq_mul_self_iff, eq_sub_iff_add_eq, or_assoc] + simp_rw [y, ← hv, ← hw, pow_two, mul_self_eq_mul_self_iff, eq_sub_iff_add_eq, or_assoc] end Quartic diff --git a/Archive/ZagierTwoSquares.lean b/Archive/ZagierTwoSquares.lean index f4846d4fea64c..0c648c6d3f851 100644 --- a/Archive/ZagierTwoSquares.lean +++ b/Archive/ZagierTwoSquares.lean @@ -198,5 +198,5 @@ theorem Nat.Prime.sq_add_sq' {p : ℕ} [h : Fact p.Prime] (hp : p % 4 = 1) : (Equiv.Perm.card_fixedPoints_modEq (p := 2) (n := 1) (complexInvo_sq k)) contrapose key rw [Set.not_nonempty_iff_eq_empty] at key - simp_rw [key, Fintype.card_eq_zero, card_fixedPoints_eq_one] + simp_rw [k, key, Fintype.card_eq_zero, card_fixedPoints_eq_one] decide diff --git a/Counterexamples/DirectSumIsInternal.lean b/Counterexamples/DirectSumIsInternal.lean index 7632fb6414a31..0c65641d433f0 100644 --- a/Counterexamples/DirectSumIsInternal.lean +++ b/Counterexamples/DirectSumIsInternal.lean @@ -85,7 +85,7 @@ theorem withSign.not_injective : apply zero_ne_one h.symm apply hinj.ne this rw [LinearMap.map_zero, LinearMap.map_add, DirectSum.toModule_lof, DirectSum.toModule_lof] - simp + simp [p1, n1] /-- And so they do not represent an internal direct sum. -/ theorem withSign.not_internal : ¬DirectSum.IsInternal withSign := diff --git a/Counterexamples/Phillips.lean b/Counterexamples/Phillips.lean index 9b2792ab43644..780bdee8c3695 100644 --- a/Counterexamples/Phillips.lean +++ b/Counterexamples/Phillips.lean @@ -287,15 +287,15 @@ theorem exists_discrete_support_nonpos (f : BoundedAdditiveMeasure α) : convert hF (s n) u using 2 · dsimp ext x - simp only [not_exists, mem_iUnion, mem_diff] + simp only [u, not_exists, mem_iUnion, mem_diff] tauto · congr 1 - simp only [s, Function.iterate_succ', Subtype.coe_mk, union_diff_left, Function.comp] + simp only [G, s, Function.iterate_succ', Subtype.coe_mk, union_diff_left, Function.comp] have I2 : ∀ n : ℕ, (n : ℝ) * (ε / 2) ≤ f ↑(s n) := by intro n induction n with | zero => - simp only [s, BoundedAdditiveMeasure.empty, id, Nat.cast_zero, zero_mul, + simp only [s, empty, BoundedAdditiveMeasure.empty, id, Nat.cast_zero, zero_mul, Function.iterate_zero, Subtype.coe_mk, le_rfl] | succ n IH => have : (s (n + 1)).1 = (s (n + 1)).1 \ (s n).1 ∪ (s n).1 := by diff --git a/Mathlib/Algebra/CharZero/Quotient.lean b/Mathlib/Algebra/CharZero/Quotient.lean index d3872fea29c7e..81c0f846bb34a 100644 --- a/Mathlib/Algebra/CharZero/Quotient.lean +++ b/Mathlib/Algebra/CharZero/Quotient.lean @@ -58,7 +58,7 @@ theorem zmultiples_zsmul_eq_zsmul_iff {ψ θ : R ⧸ AddSubgroup.zmultiples p} { -- Porting note: Introduced Zp notation to shorten lines let Zp := AddSubgroup.zmultiples p have : (Quotient.mk _ : R → R ⧸ Zp) = ((↑) : R → R ⧸ Zp) := rfl - simp only [this] + simp only [Zp, this] simp_rw [← QuotientAddGroup.mk_zsmul, ← QuotientAddGroup.mk_add, QuotientAddGroup.eq_iff_sub_mem, ← smul_sub, ← sub_sub] exact AddSubgroup.zsmul_mem_zmultiples_iff_exists_sub_div hz diff --git a/Mathlib/Algebra/ContinuedFractions/Computation/Approximations.lean b/Mathlib/Algebra/ContinuedFractions/Computation/Approximations.lean index e2e525d76dcdc..079913152e68b 100644 --- a/Mathlib/Algebra/ContinuedFractions/Computation/Approximations.lean +++ b/Mathlib/Algebra/ContinuedFractions/Computation/Approximations.lean @@ -219,7 +219,7 @@ theorem fib_le_of_contsAux_b : -- use the recurrence of `contsAux` simp only [Nat.succ_eq_add_one, Nat.add_assoc, Nat.reduceAdd] suffices (fib n : K) + fib (n + 1) ≤ gp.a * ppconts.b + gp.b * pconts.b by - simpa [fib_add_two, add_comm, contsAux_recurrence s_ppred_nth_eq ppconts_eq pconts_eq] + simpa [g, fib_add_two, add_comm, contsAux_recurrence s_ppred_nth_eq ppconts_eq pconts_eq] -- make use of the fact that `gp.a = 1` suffices (fib n : K) + fib (n + 1) ≤ ppconts.b + gp.b * pconts.b by simpa [of_partNum_eq_one <| partNum_eq_s_a s_ppred_nth_eq] @@ -267,7 +267,7 @@ theorem zero_le_of_contsAux_b : 0 ≤ ((of v).contsAux n).b := by · simp [zero_le_one] · have : g.contsAux (n + 2) = g.contsAux (n + 1) := contsAux_stable_step_of_terminated terminated - simp only [this, IH] + simp only [g, this, IH] · -- non-terminating case calc (0 : K) ≤ fib (n + 1) := mod_cast (n + 1).fib.zero_le diff --git a/Mathlib/Algebra/ContinuedFractions/Computation/CorrectnessTerminating.lean b/Mathlib/Algebra/ContinuedFractions/Computation/CorrectnessTerminating.lean index a0117d970346f..fea70d66955d1 100644 --- a/Mathlib/Algebra/ContinuedFractions/Computation/CorrectnessTerminating.lean +++ b/Mathlib/Algebra/ContinuedFractions/Computation/CorrectnessTerminating.lean @@ -199,7 +199,7 @@ theorem compExactValue_correctness_of_stream_eq_some : nextConts, nextNum, nextDen] have hfr : (IntFractPair.of (1 / ifp_n.fr)).fr = f := rfl rw [one_div, if_neg _, ← one_div, hfr] - · field_simp [hA, hB] + · field_simp [pA, pB, ppA, ppB, pconts, ppconts, hA, hB] ac_rfl · rwa [inv_eq_one_div, hfr] diff --git a/Mathlib/Algebra/ContinuedFractions/Computation/TerminatesIffRat.lean b/Mathlib/Algebra/ContinuedFractions/Computation/TerminatesIffRat.lean index 29595521c11a7..2a78678ad1a4a 100644 --- a/Mathlib/Algebra/ContinuedFractions/Computation/TerminatesIffRat.lean +++ b/Mathlib/Algebra/ContinuedFractions/Computation/TerminatesIffRat.lean @@ -80,7 +80,7 @@ nonrec theorem exists_gcf_pair_rat_eq_of_nth_contsAux : · use pred_conts have : g.contsAux (n + 2) = g.contsAux (n + 1) := contsAux_stable_of_terminated (n + 1).le_succ s_ppred_nth_eq - simp only [this, pred_conts_eq] + simp only [g, this, pred_conts_eq] -- option.some · -- invoke the IH a second time obtain ⟨ppred_conts, ppred_conts_eq⟩ := @@ -88,7 +88,7 @@ nonrec theorem exists_gcf_pair_rat_eq_of_nth_contsAux : obtain ⟨a_eq_one, z, b_eq_z⟩ : gp_n.a = 1 ∧ ∃ z : ℤ, gp_n.b = (z : K) := of_partNum_eq_one_and_exists_int_partDen_eq s_ppred_nth_eq -- finally, unfold the recurrence to obtain the required rational value. - simp only [a_eq_one, b_eq_z, + simp only [g, a_eq_one, b_eq_z, contsAux_recurrence s_ppred_nth_eq ppred_conts_eq pred_conts_eq] use nextConts 1 (z : ℚ) ppred_conts pred_conts cases ppred_conts; cases pred_conts diff --git a/Mathlib/Algebra/ContinuedFractions/ConvergentsEquiv.lean b/Mathlib/Algebra/ContinuedFractions/ConvergentsEquiv.lean index fa5e59f20d0ad..80a89dc903ff4 100644 --- a/Mathlib/Algebra/ContinuedFractions/ConvergentsEquiv.lean +++ b/Mathlib/Algebra/ContinuedFractions/ConvergentsEquiv.lean @@ -303,7 +303,7 @@ theorem succ_nth_conv_eq_squashGCF_nth_conv [Field K] ((pb + a / b) * pA + pa * ppA) / ((pb + a / b) * pB + pa * ppB) = (b * (pb * pA + pa * ppA) + a * pA) / (b * (pb * pB + pa * ppB) + a * pB) by obtain ⟨eq1, eq2, eq3, eq4⟩ : pA' = pA ∧ pB' = pB ∧ ppA' = ppA ∧ ppB' = ppB := by - simp [*, pA', pB', ppA', ppB', + simp [*, g', pA, pB, ppA, ppB, pA', pB', ppA', ppB', (contsAux_eq_contsAux_squashGCF_of_le <| le_refl <| n' + 1).symm, (contsAux_eq_contsAux_squashGCF_of_le n'.le_succ).symm] symm diff --git a/Mathlib/Algebra/GroupWithZero/WithZero.lean b/Mathlib/Algebra/GroupWithZero/WithZero.lean index 66c0edd8d8c17..8acf0b8bcb620 100644 --- a/Mathlib/Algebra/GroupWithZero/WithZero.lean +++ b/Mathlib/Algebra/GroupWithZero/WithZero.lean @@ -255,7 +255,10 @@ instance groupWithZero : GroupWithZero (WithZero α) where __ := divInvMonoid __ := nontrivial inv_zero := WithZero.inv_zero - mul_inv_cancel a ha := by lift a to α using ha; norm_cast; apply mul_inv_cancel + mul_inv_cancel a ha := by + lift a to α using ha + norm_cast + apply mul_inv_cancel /-- Any group is isomorphic to the units of itself adjoined with `0`. -/ diff --git a/Mathlib/Algebra/Homology/ShortComplex/Exact.lean b/Mathlib/Algebra/Homology/ShortComplex/Exact.lean index 6c3f809ed38d6..4e407ab86ecdf 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/Exact.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/Exact.lean @@ -930,14 +930,14 @@ variable (F : C ⥤ D) [Preadditive C] [Preadditive D] [HasZeroObject C] instance : F.PreservesMonomorphisms where preserves {X Y} f hf := by let S := ShortComplex.mk (0 : X ⟶ X) f zero_comp - exact ((S.map F).exact_iff_mono (by simp)).1 + exact ((S.map F).exact_iff_mono (by simp [S])).1 (((S.exact_iff_mono rfl).2 hf).map F) instance : F.PreservesEpimorphisms where preserves {X Y} f hf := by let S := ShortComplex.mk f (0 : Y ⟶ Y) comp_zero - exact ((S.map F).exact_iff_epi (by simp)).1 + exact ((S.map F).exact_iff_epi (by simp [S])).1 (((S.exact_iff_epi rfl).2 hf).map F) diff --git a/Mathlib/Algebra/Lie/TraceForm.lean b/Mathlib/Algebra/Lie/TraceForm.lean index 55ee3be037f29..ff980b97b8cde 100644 --- a/Mathlib/Algebra/Lie/TraceForm.lean +++ b/Mathlib/Algebra/Lie/TraceForm.lean @@ -275,7 +275,8 @@ lemma lowerCentralSeries_one_inf_center_le_ker_traceForm [Module.Free R M] [Modu replace hzc : 1 ⊗ₜ[R] z ∈ LieAlgebra.center A (A ⊗[R] L) := by simp only [mem_maxTrivSubmodule] at hzc ⊢ intro y - exact y.induction_on rfl (fun a u ↦ by simp [hzc u]) (fun u v hu hv ↦ by simp [hu, hv]) + exact y.induction_on rfl (fun a u ↦ by simp [hzc u]) + (fun u v hu hv ↦ by simp [A, hu, hv]) apply LinearMap.trace_comp_eq_zero_of_commute_of_trace_restrict_eq_zero · exact IsTriangularizable.maxGenEigenspace_eq_top (1 ⊗ₜ[R] x) · exact fun μ ↦ trace_toEnd_eq_zero_of_mem_lcs A (A ⊗[R] L) diff --git a/Mathlib/Algebra/Lie/Weights/Basic.lean b/Mathlib/Algebra/Lie/Weights/Basic.lean index 8ddf44e6372c7..f3bb8d4eb7128 100644 --- a/Mathlib/Algebra/Lie/Weights/Basic.lean +++ b/Mathlib/Algebra/Lie/Weights/Basic.lean @@ -414,7 +414,7 @@ def posFittingCompOf (x : L) : LieSubmodule R L M := obtain ⟨q, hq⟩ := h₁.add_pow_dvd_pow_of_pow_eq_zero_right (N + k).le_succ hN use toModuleHom R L M (q (y ⊗ₜ m)) change (φ ^ k).comp ((toModuleHom R L M : L ⊗[R] M →ₗ[R] M)) _ = _ - simp [φ, f₁, f₂, LinearMap.commute_pow_left_of_commute h₂, + simp [φ, f₁, f₂, LinearMap.commute_pow_left_of_commute h₂, LinearMap.comp_apply (g := (f₁ + f₂) ^ k), ← LinearMap.comp_apply (g := q), ← LinearMap.mul_eq_comp, ← hq] } diff --git a/Mathlib/Algebra/Module/FreeLocus.lean b/Mathlib/Algebra/Module/FreeLocus.lean index 770e2b8cab031..8389295b32ac8 100644 --- a/Mathlib/Algebra/Module/FreeLocus.lean +++ b/Mathlib/Algebra/Module/FreeLocus.lean @@ -217,7 +217,7 @@ lemma isLocallyConstant_rankAtStalk_freeLocus [Module.FinitePresentation R M] : have : IsLocalizedModule (Algebra.algebraMapSubmonoid _ p.asIdeal.primeCompl) l := IsLocalizedModule.of_restrictScalars p.asIdeal.primeCompl .. have := Module.finrank_of_isLocalizedModule_of_free Rₚ p' l - simp [rankAtStalk, this, hf''] + simp [Rₚ, rankAtStalk, this, hf''] lemma isLocallyConstant_rankAtStalk [Module.FinitePresentation R M] [Module.Flat R M] : IsLocallyConstant (rankAtStalk (R := R) M) := by diff --git a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean index 4306631e05eb2..952782f698cdd 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean @@ -527,7 +527,7 @@ private lemma respects_isOpenImmersion_aux let e : (U.ι ⁻¹ᵁ s).toScheme ≅ s := IsOpenImmersion.isoOfRangeEq ((U.ι ⁻¹ᵁ s).ι ≫ U.ι) s.1.ι (by simpa [Set.range_comp, Set.image_preimage_eq_iff, heq] using le_sSup s.2) have heq : (V s).ι ≫ f ≫ U.ι = f' ≫ e.hom ≫ s.1.ι := by - simp only [IsOpenImmersion.isoOfRangeEq_hom_fac, f', e, morphismRestrict_ι_assoc] + simp only [V, IsOpenImmersion.isoOfRangeEq_hom_fac, f', e, morphismRestrict_ι_assoc] rw [heq, ← Category.assoc] refine this _ ?_ ?_ · rwa [P.cancel_right_of_respectsIso] diff --git a/Mathlib/Analysis/Convex/Jensen.lean b/Mathlib/Analysis/Convex/Jensen.lean index 332967af25195..0fb76f66c1e58 100644 --- a/Mathlib/Analysis/Convex/Jensen.lean +++ b/Mathlib/Analysis/Convex/Jensen.lean @@ -117,7 +117,7 @@ lemma StrictConvexOn.map_sum_lt (hf : StrictConvexOn 𝕜 s f) (h₀ : ∀ i ∈ have := h₀ j <| by simp have := h₀ k <| by simp let c := w j + w k - have hc : w j / c + w k / c = 1 := by field_simp + have hc : w j / c + w k / c = 1 := by field_simp [c] 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 congrm f ?_ diff --git a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean index 660294ee97883..616027a97b9a7 100644 --- a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean +++ b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean @@ -395,7 +395,7 @@ theorem lintegral_pow_le_pow_lintegral_fderiv {u : E → F} have h2c : μ = c • ((volume : Measure (ι → ℝ)).map e.symm) := isAddLeftInvariant_eq_smul .. have h3c : (c : ℝ≥0∞) ≠ 0 := by simp_rw [ne_eq, ENNReal.coe_eq_zero, hc.ne', not_false_eq_true] have h0C : C = (c * ‖(e.symm : (ι → ℝ) →L[ℝ] E)‖₊ ^ p) * (c ^ p)⁻¹ := by - simp_rw [C, lintegralPowLePowLIntegralFDerivConst] + simp_rw [c, ι, C, e, lintegralPowLePowLIntegralFDerivConst] have hC : C * c ^ p = c * ‖(e.symm : (ι → ℝ) →L[ℝ] E)‖₊ ^ p := by rw [h0C, inv_mul_cancel_right₀ (NNReal.rpow_pos hc).ne'] rw [h2c, ENNReal.smul_def, lintegral_smul_measure, lintegral_smul_measure] @@ -506,7 +506,7 @@ theorem eLpNorm_le_eLpNorm_fderiv_of_eq_inner {u : E → F'} rw [← inv_inj, hp'] field_simp [n', NNReal.conjExponent] · norm_cast - simp_rw [eLpNormLESNormFDerivOfEqInnerConst] + simp_rw [n', n, eLpNormLESNormFDerivOfEqInnerConst] field_simp -- the case `p > 1` let q := Real.conjExponent p @@ -570,7 +570,7 @@ theorem eLpNorm_le_eLpNorm_fderiv_of_eq_inner {u : E → F'} gcongr convert ENNReal.lintegral_mul_le_Lp_mul_Lq μ (.symm <| .conjExponent <| show 1 < (p : ℝ) from hp) ?_ ?_ using 5 - · simp_rw [← ENNReal.rpow_mul, ← h3γ] + · simp [γ, n, q, ← ENNReal.rpow_mul, ← h3γ] · borelize F' fun_prop · fun_prop @@ -646,7 +646,7 @@ theorem eLpNorm_le_eLpNorm_fderiv_of_eq [FiniteDimensional ℝ F] _ ≤ C₁ * C * (C₂ * eLpNorm (fderiv ℝ u) p μ) := by gcongr; exact eLpNorm_le_nnreal_smul_eLpNorm_of_ae_le_mul (Eventually.of_forall h4v) p _ = SNormLESNormFDerivOfEqConst F μ p * eLpNorm (fderiv ℝ u) p μ := by - simp_rw [SNormLESNormFDerivOfEqConst] + simp_rw [C₂, C₁, C, e, SNormLESNormFDerivOfEqConst] push_cast simp_rw [mul_assoc] diff --git a/Mathlib/Analysis/InnerProductSpace/PiL2.lean b/Mathlib/Analysis/InnerProductSpace/PiL2.lean index 4813828390c93..15e7d1d6c28ab 100644 --- a/Mathlib/Analysis/InnerProductSpace/PiL2.lean +++ b/Mathlib/Analysis/InnerProductSpace/PiL2.lean @@ -953,7 +953,7 @@ noncomputable def LinearIsometry.extend (L : S →ₗᵢ[𝕜] V) : V →ₗᵢ[ have Lp1x : L (p1 x) ∈ LinearMap.range L.toLinearMap := LinearMap.mem_range_self L.toLinearMap (p1 x) have Lp2x : L3 (p2 x) ∈ (LinearMap.range L.toLinearMap)ᗮ := by - simp only [LinearIsometry.coe_comp, Function.comp_apply, Submodule.coe_subtypeₗᵢ, + simp only [LS, LinearIsometry.coe_comp, Function.comp_apply, Submodule.coe_subtypeₗᵢ, ← Submodule.range_subtype LSᗮ] apply LinearMap.mem_range_self apply Submodule.inner_right_of_mem_orthogonal Lp1x Lp2x diff --git a/Mathlib/Analysis/InnerProductSpace/TwoDim.lean b/Mathlib/Analysis/InnerProductSpace/TwoDim.lean index afdbc274d603b..18f8a99312e5f 100644 --- a/Mathlib/Analysis/InnerProductSpace/TwoDim.lean +++ b/Mathlib/Analysis/InnerProductSpace/TwoDim.lean @@ -540,10 +540,10 @@ attribute [local instance] Complex.finrank_real_complex_fact @[simp] protected theorem areaForm (w z : ℂ) : Complex.orientation.areaForm w z = (conj w * z).im := by let o := Complex.orientation - simp only [o.areaForm_to_volumeForm, o.volumeForm_robust Complex.orthonormalBasisOneI rfl, - Basis.det_apply, Matrix.det_fin_two, Basis.toMatrix_apply, toBasis_orthonormalBasisOneI, - Matrix.cons_val_zero, coe_basisOneI_repr, Matrix.cons_val_one, Matrix.head_cons, mul_im, - conj_re, conj_im] + simp only [o, o.areaForm_to_volumeForm, + o.volumeForm_robust Complex.orthonormalBasisOneI rfl, Basis.det_apply, Matrix.det_fin_two, + Basis.toMatrix_apply, toBasis_orthonormalBasisOneI, Matrix.cons_val_zero, coe_basisOneI_repr, + Matrix.cons_val_one, Matrix.head_cons, mul_im, conj_re, conj_im] ring @[simp] diff --git a/Mathlib/CategoryTheory/Functor/Flat.lean b/Mathlib/CategoryTheory/Functor/Flat.lean index 2d0444d695f18..0cb417530c71d 100644 --- a/Mathlib/CategoryTheory/Functor/Flat.lean +++ b/Mathlib/CategoryTheory/Functor/Flat.lean @@ -82,6 +82,7 @@ theorem RepresentablyFlat.id : RepresentablyFlat (𝟭 C) := inferInstance theorem RepresentablyCoflat.id : RepresentablyCoflat (𝟭 C) := inferInstance +set_option maxHeartbeats 400000 in instance RepresentablyFlat.comp (G : D ⥤ E) [RepresentablyFlat F] [RepresentablyFlat G] : RepresentablyFlat (F ⋙ G) := by refine ⟨fun X => IsCofiltered.of_cone_nonempty.{0} _ (fun {J} _ _ H => ?_)⟩ @@ -91,6 +92,7 @@ instance RepresentablyFlat.comp (G : D ⥤ E) [RepresentablyFlat F] map := fun {j j'} f => StructuredArrow.homMk (H.map f).right (congrArg CommaMorphism.right (c₁.w f)) } obtain ⟨c₂⟩ := IsCofiltered.cone_nonempty H₂ + simp only [H₂] at c₂ exact ⟨⟨StructuredArrow.mk (c₁.pt.hom ≫ G.map c₂.pt.hom), ⟨fun j => StructuredArrow.homMk (c₂.π.app j).right (by simp [← G.map_comp, (c₂.π.app j).w]), fun j j' f => by simpa using (c₂.w f).symm⟩⟩⟩ diff --git a/Mathlib/CategoryTheory/Sites/DenseSubsite/SheafEquiv.lean b/Mathlib/CategoryTheory/Sites/DenseSubsite/SheafEquiv.lean index 487e2c24802c2..ecb6633e6f906 100644 --- a/Mathlib/CategoryTheory/Sites/DenseSubsite/SheafEquiv.lean +++ b/Mathlib/CategoryTheory/Sites/DenseSubsite/SheafEquiv.lean @@ -91,7 +91,7 @@ lemma isIso_ranCounit_app_of_isDenseSubsite (Y : Sheaf J A) (U X) : RightExtension.coneAt_pt, RightExtension.mk_left, RightExtension.coneAt_π_app, const_obj_obj, op_obj, StructuredArrow.mk_hom_eq_self, map_id, whiskeringLeft_obj_obj, RightExtension.mk_hom, Category.id_comp, StructuredArrow.mk_left, unop_id] at this - simp only [id_obj, yoneda_map_app, this] + simp only [c, id_obj, yoneda_map_app, this] apply Y.2.hom_ext ⟨_, IsDenseSubsite.imageSieve_mem J K G (𝟙 (G.obj U))⟩ _ _ fun I ↦ ?_ apply (Y.2 X _ (IsDenseSubsite.equalizer_mem J K G (l _ _ _ _ _ I.hf) I.f (by simp [hl]))).isSeparatedFor.ext fun V iUV (hiUV : _ = _) ↦ ?_ diff --git a/Mathlib/Combinatorics/Colex.lean b/Mathlib/Combinatorics/Colex.lean index 98e0a7d8d89ae..4b9156cab8f08 100644 --- a/Mathlib/Combinatorics/Colex.lean +++ b/Mathlib/Combinatorics/Colex.lean @@ -318,7 +318,8 @@ lemma toColex_le_toColex_iff_max'_mem : refine ⟨fun h hst ↦ ?_, fun h a has hat ↦ ?_⟩ · set m := (s ∆ t).max' (symmDiff_nonempty.2 hst) by_contra hmt - have hms : m ∈ s := by simpa [mem_symmDiff, hmt] using max'_mem _ <| symmDiff_nonempty.2 hst + have hms : m ∈ s := by + simpa [m, mem_symmDiff, hmt] using max'_mem _ <| symmDiff_nonempty.2 hst have ⟨b, hbt, hbs, hmb⟩ := h hms hmt exact lt_irrefl _ <| (max'_lt_iff _ _).1 (hmb.lt_of_ne <| ne_of_mem_of_not_mem hms hbs) _ <| mem_symmDiff.2 <| Or.inr ⟨hbt, hbs⟩ diff --git a/Mathlib/Combinatorics/Hall/Basic.lean b/Mathlib/Combinatorics/Hall/Basic.lean index 1cd190b29bab4..cfb20a097cdc9 100644 --- a/Mathlib/Combinatorics/Hall/Basic.lean +++ b/Mathlib/Combinatorics/Hall/Basic.lean @@ -184,11 +184,8 @@ theorem Fintype.all_card_le_rel_image_card_iff_exists_injective {α : Type u} { rw [← Set.toFinset_card] apply congr_arg ext b - -- Porting note: added `Set.mem_toFinset` (fixed by https://github.com/leanprover/lean4/pull/6123?) - simp [r', Rel.image, (Set.mem_toFinset)] - -- Porting note: added `Set.mem_toFinset` (fixed by https://github.com/leanprover/lean4/pull/6123?) - have h' : ∀ (f : α → β) (x), r x (f x) ↔ f x ∈ r' x := by - simp [r', Rel.image, (Set.mem_toFinset)] + simp [r', Rel.image] + have h' : ∀ (f : α → β) (x), r x (f x) ↔ f x ∈ r' x := by simp [r', Rel.image] simp only [h, h'] apply Finset.all_card_le_biUnion_card_iff_exists_injective diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Equitabilise.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Equitabilise.lean index 30e75e1ec0e80..b82ae053e2ea0 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Equitabilise.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Equitabilise.lean @@ -89,7 +89,7 @@ theorem equitabilise_aux (hs : a * m + b * (m + 1) = #s) : · simp only [extend_parts, mem_insert, forall_eq_or_imp, and_iff_left hR₁, htn, hn] exact ite_eq_or_eq _ _ _ · exact fun x hx => (card_le_card sdiff_subset).trans (Nat.lt_succ_iff.1 <| h _ hx) - simp_rw [extend_parts, filter_insert, htn, m.succ_ne_self.symm.ite_eq_right_iff] + simp_rw [extend_parts, filter_insert, htn, n, m.succ_ne_self.symm.ite_eq_right_iff] split_ifs with ha · rw [hR₃, if_pos ha] rw [card_insert_of_not_mem, hR₃, if_neg ha, tsub_add_cancel_of_le] diff --git a/Mathlib/Computability/PartrecCode.lean b/Mathlib/Computability/PartrecCode.lean index 02daf224ca8e5..52df16d8893d0 100644 --- a/Mathlib/Computability/PartrecCode.lean +++ b/Mathlib/Computability/PartrecCode.lean @@ -164,7 +164,7 @@ private theorem encode_ofNatCode : ∀ n, encodeCode (ofNatCode n) = n conv_rhs => rw [← Nat.bit_decomp n, ← Nat.bit_decomp n.div2] simp only [ofNatCode.eq_5] cases n.bodd <;> cases n.div2.bodd <;> - simp [encodeCode, ofNatCode, IH, IH1, IH2, Nat.bit_val] + simp [m, encodeCode, ofNatCode, IH, IH1, IH2, Nat.bit_val] instance instDenumerable : Denumerable Code := mk' @@ -325,7 +325,7 @@ theorem rec_prim' {α σ} [Primcodable α] [Primcodable σ] {c : α → Code} (h (Nat.succ_le_succ (Nat.le_add_right ..)) have m1 : m.unpair.1 < n + 4 := lt_of_le_of_lt m.unpair_left_le hm have m2 : m.unpair.2 < n + 4 := lt_of_le_of_lt m.unpair_right_le hm - simp [G₁]; simp [m, List.getElem?_map, List.getElem?_range, hm, m1, m2] + simp [G₁, m, List.getElem?_map, List.getElem?_range, hm, m1, m2] rw [show ofNat Code (n + 4) = ofNatCode (n + 4) from rfl] simp [ofNatCode] cases n.bodd <;> cases n.div2.bodd <;> rfl @@ -435,7 +435,7 @@ theorem rec_computable {α σ} [Primcodable α] [Primcodable σ] {c : α → Cod (Nat.succ_le_succ (Nat.le_add_right ..)) have m1 : m.unpair.1 < n + 4 := lt_of_le_of_lt m.unpair_left_le hm have m2 : m.unpair.2 < n + 4 := lt_of_le_of_lt m.unpair_right_le hm - simp [G₁]; simp [m, List.getElem?_map, List.getElem?_range, hm, m1, m2] + simp [G₁, m, List.getElem?_map, List.getElem?_range, hm, m1, m2] rw [show ofNat Code (n + 4) = ofNatCode (n + 4) from rfl] simp [ofNatCode] cases n.bodd <;> cases n.div2.bodd <;> rfl @@ -948,7 +948,7 @@ theorem evaln_prim : Primrec fun a : (ℕ × Code) × ℕ => evaln a.1.1 a.1.2 a rw [hg (Nat.pair_lt_pair_right _ lg)] cases evaln k cg n · rfl - simp [hg (Nat.pair_lt_pair_right _ lf)] + simp [k, hg (Nat.pair_lt_pair_right _ lf)] · cases' encode_lt_prec cf cg with lf lg rw [hg (Nat.pair_lt_pair_right _ lf)] cases n.unpair.2 @@ -957,7 +957,7 @@ theorem evaln_prim : Primrec fun a : (ℕ × Code) × ℕ => evaln a.1.1 a.1.2 a rw [hg (Nat.pair_lt_pair_left _ k'.lt_succ_self)] cases evaln k' _ _ · rfl - simp [hg (Nat.pair_lt_pair_right _ lg)] + simp [k, hg (Nat.pair_lt_pair_right _ lg)] · have lf := encode_lt_rfind' cf rw [hg (Nat.pair_lt_pair_right _ lf)] cases' evaln k cf n with x @@ -1005,7 +1005,7 @@ theorem fixed_point {f : Code → Code} (hf : Computable f) : ∃ c : Code, eval ⟨curry cg (encode cF), funext fun n => show eval (f (curry cg (encode cF))) n = eval (curry cg (encode cF)) n by - simp [g, eg', eF', Part.map_id']⟩ + simp [F, g, eg', eF', Part.map_id']⟩ theorem fixed_point₂ {f : Code → ℕ →. ℕ} (hf : Partrec₂ f) : ∃ c : Code, eval c = f c := let ⟨cf, ef⟩ := exists_code.1 hf diff --git a/Mathlib/Computability/Primrec.lean b/Mathlib/Computability/Primrec.lean index 76aceca0eb52f..1eaa637c25e5e 100644 --- a/Mathlib/Computability/Primrec.lean +++ b/Mathlib/Computability/Primrec.lean @@ -762,7 +762,7 @@ private theorem list_foldl' {f : α → List β} {g : α → σ} {h : α → σ | zero => rfl | succ n IH => simp only [iterate_succ, comp_apply] - cases' l with b l <;> simp [IH] + cases' l with b l <;> simp [G, IH] private theorem list_cons' : (haveI := prim H; Primrec₂ (@List.cons β)) := letI := prim H diff --git a/Mathlib/Data/ENat/Basic.lean b/Mathlib/Data/ENat/Basic.lean index 5b96e88c13f2d..c3d1a69c8162c 100644 --- a/Mathlib/Data/ENat/Basic.lean +++ b/Mathlib/Data/ENat/Basic.lean @@ -339,7 +339,9 @@ lemma add_one_natCast_le_withTop_of_lt {m : ℕ} {n : WithTop ℕ∞} (h : m < n match n with | ⊤ => exact Iff.rfl | (⊤ : ℕ∞) => simp - | (n : ℕ) => norm_cast; simp only [coe_ne_top, iff_false, ne_eq] + | (n : ℕ) => + norm_cast + simp only [coe_ne_top, iff_false, ne_eq] @[simp] lemma natCast_ne_coe_top (n : ℕ) : (n : WithTop ℕ∞) ≠ (⊤ : ℕ∞) := nofun diff --git a/Mathlib/Data/Matrix/Rank.lean b/Mathlib/Data/Matrix/Rank.lean index 13db0d03e2bc9..69228715dac42 100644 --- a/Mathlib/Data/Matrix/Rank.lean +++ b/Mathlib/Data/Matrix/Rank.lean @@ -142,8 +142,8 @@ theorem rank_eq_finrank_range_toLin [Finite m] [DecidableEq n] {M₁ M₂ : Type have aux₂ := Basis.equiv_apply (Pi.basisFun R n) i v₂ rw [toLin_eq_toLin', toLin'_apply'] at aux₁ rw [Pi.basisFun_apply] at aux₁ aux₂ - simp only [e₁, e₁, LinearMap.comp_apply, LinearEquiv.coe_coe, Equiv.refl_apply, aux₁, aux₂, - LinearMap.coe_single, toLin_self, map_sum, LinearEquiv.map_smul, Basis.equiv_apply] + simp only [e₁, e₂, LinearMap.comp_apply, LinearEquiv.coe_coe, Equiv.refl_apply, + aux₁, aux₂, LinearMap.coe_single, toLin_self, map_sum, LinearEquiv.map_smul, Basis.equiv_apply] theorem rank_le_card_height [Fintype m] [StrongRankCondition R] (A : Matrix m n R) : A.rank ≤ Fintype.card m := by diff --git a/Mathlib/Data/Nat/Fib/Zeckendorf.lean b/Mathlib/Data/Nat/Fib/Zeckendorf.lean index ea2b791ad22c0..8e0217a7dc2f1 100644 --- a/Mathlib/Data/Nat/Fib/Zeckendorf.lean +++ b/Mathlib/Data/Nat/Fib/Zeckendorf.lean @@ -118,7 +118,7 @@ Note: For unfolding, you should use the equational lemmas `Nat.zeckendorf_zero` def zeckendorf : ℕ → List ℕ | 0 => [] | m@(_ + 1) => - let a := greatestFib m + letI a := greatestFib m a :: zeckendorf (m - fib a) decreasing_by simp_wf; subst_vars; apply zeckendorf_aux (zero_lt_succ _) diff --git a/Mathlib/Data/Nat/ModEq.lean b/Mathlib/Data/Nat/ModEq.lean index 29f36306ea23a..ebe5c44a758a4 100644 --- a/Mathlib/Data/Nat/ModEq.lean +++ b/Mathlib/Data/Nat/ModEq.lean @@ -244,8 +244,8 @@ lemma cancel_left_div_gcd (hm : 0 < m) (h : c * a ≡ c * b [MOD m]) : a ≡ b 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, - Nat.div_self (gcd_pos_of_pos_left c hm)] + simp only [d, ← Int.natCast_div, Int.gcd_natCast_natCast (m / d) (c / d), + gcd_div hmd hcd, Nat.div_self (gcd_pos_of_pos_left c hm)] /-- To cancel a common factor `c` from a `ModEq` we must divide the modulus `m` by `gcd m c` -/ lemma cancel_right_div_gcd (hm : 0 < m) (h : a * c ≡ b * c [MOD m]) : a ≡ b [MOD m / gcd m c] := by diff --git a/Mathlib/Data/Nat/Pairing.lean b/Mathlib/Data/Nat/Pairing.lean index bda4acd194b8d..66a84f6f6f69a 100644 --- a/Mathlib/Data/Nat/Pairing.lean +++ b/Mathlib/Data/Nat/Pairing.lean @@ -45,10 +45,10 @@ theorem pair_unpair (n : ℕ) : pair (unpair n).1 (unpair n).2 = n := by dsimp only [unpair]; let s := sqrt n have sm : s * s + (n - s * s) = n := Nat.add_sub_cancel' (sqrt_le _) split_ifs with h - · simp [pair, h, sm] + · simp [s, pair, h, sm] · have hl : n - s * s - s ≤ s := Nat.sub_le_iff_le_add.2 (Nat.sub_le_iff_le_add'.2 <| by rw [← Nat.add_assoc]; apply sqrt_le_add) - simp [pair, hl.not_lt, Nat.add_assoc, Nat.add_sub_cancel' (le_of_not_gt h), sm] + simp [s, pair, hl.not_lt, Nat.add_assoc, Nat.add_sub_cancel' (le_of_not_gt h), sm] theorem pair_unpair' {n a b} (H : unpair n = (a, b)) : pair a b = n := by simpa [H] using pair_unpair n @@ -80,7 +80,7 @@ theorem pair_eq_pair {a b c d : ℕ} : pair a b = pair c d ↔ a = c ∧ b = d : theorem unpair_lt {n : ℕ} (n1 : 1 ≤ n) : (unpair n).1 < n := by let s := sqrt n simp only [unpair, Nat.sub_le_iff_le_add] - by_cases h : n - s * s < s <;> simp only [h, ↓reduceIte] + by_cases h : n - s * s < s <;> simp [s, h, ↓reduceIte] · exact lt_of_lt_of_le h (sqrt_le_self _) · simp only [not_lt] at h have s0 : 0 < s := sqrt_pos.2 n1 diff --git a/Mathlib/Data/Nat/Prime/Basic.lean b/Mathlib/Data/Nat/Prime/Basic.lean index 68d16200182d3..e1452b5e145a5 100644 --- a/Mathlib/Data/Nat/Prime/Basic.lean +++ b/Mathlib/Data/Nat/Prime/Basic.lean @@ -216,8 +216,8 @@ theorem ne_one_iff_exists_prime_dvd : ∀ {n}, n ≠ 1 ↔ ∃ p : ℕ, p.Prime | 1 => by simp [Nat.not_prime_one] | n + 2 => by let a := n + 2 - let ha : a ≠ 1 := Nat.succ_succ_ne_one n - simp only [true_iff, Ne, not_false_iff, ha] + have ha : a ≠ 1 := Nat.succ_succ_ne_one n + simp only [a, true_iff, Ne, not_false_iff, ha] exact ⟨a.minFac, Nat.minFac_prime ha, a.minFac_dvd⟩ theorem eq_one_iff_not_exists_prime_dvd {n : ℕ} : n = 1 ↔ ∀ p : ℕ, p.Prime → ¬p ∣ n := by diff --git a/Mathlib/Data/Nat/Prime/Defs.lean b/Mathlib/Data/Nat/Prime/Defs.lean index 30bc5d25b64e8..522dbff9c27b4 100644 --- a/Mathlib/Data/Nat/Prime/Defs.lean +++ b/Mathlib/Data/Nat/Prime/Defs.lean @@ -225,12 +225,13 @@ theorem minFacAux_has_prop {n : ℕ} (n2 : 2 ≤ n) : · have pp : Prime n := prime_def_le_sqrt.2 ⟨n2, fun m m2 l d => not_lt_of_ge l <| lt_of_lt_of_le (sqrt_lt.2 h) (a m m2 d)⟩ - simpa [h] using ⟨n2, dvd_rfl, fun m m2 d => le_of_eq ((dvd_prime_two_le pp m2).1 d).symm⟩ + simpa only [k, h] using + ⟨n2, dvd_rfl, fun m m2 d => le_of_eq ((dvd_prime_two_le pp m2).1 d).symm⟩ have k2 : 2 ≤ k := by subst e apply Nat.le_add_left - simp only [h, ↓reduceIte] - by_cases dk : k ∣ n <;> simp only [dk, ↓reduceIte] + simp only [k, h, ↓reduceIte] + by_cases dk : k ∣ n <;> simp only [k, dk, ↓reduceIte] · exact ⟨k2, dk, a⟩ · refine have := minFac_lemma n k h diff --git a/Mathlib/Data/Nat/Sqrt.lean b/Mathlib/Data/Nat/Sqrt.lean index d5ebe7765840f..09e8891acf36d 100644 --- a/Mathlib/Data/Nat/Sqrt.lean +++ b/Mathlib/Data/Nat/Sqrt.lean @@ -48,9 +48,9 @@ lemma sqrt.iter_sq_le (n guess : ℕ) : sqrt.iter n guess * sqrt.iter n guess unfold sqrt.iter let next := (guess + n / guess) / 2 if h : next < guess then - simpa only [dif_pos h] using sqrt.iter_sq_le n next + simpa only [next, dif_pos h] using sqrt.iter_sq_le n next else - simp only [dif_neg h] + simp only [next, dif_neg h] apply Nat.mul_le_of_le_div apply Nat.le_of_add_le_add_left (a := guess) rw [← Nat.mul_two, ← le_div_iff_mul_le] diff --git a/Mathlib/FieldTheory/PurelyInseparable.lean b/Mathlib/FieldTheory/PurelyInseparable.lean index 6085b263e3ab1..1d226a22d9443 100644 --- a/Mathlib/FieldTheory/PurelyInseparable.lean +++ b/Mathlib/FieldTheory/PurelyInseparable.lean @@ -959,7 +959,7 @@ theorem adjoin_eq_of_isAlgebraic [Algebra.IsAlgebraic F E] : rw [lift_top, lift_adjoin] at h haveI : IsScalarTower F S K := IsScalarTower.of_algebraMap_eq (congrFun rfl) rw [← h, ← map_eq_of_separableClosure_eq_bot F (separableClosure_eq_bot E K)] - simp only [coe_map, IsScalarTower.coe_toAlgHom', IntermediateField.algebraMap_apply] + simp only [S, coe_map, IsScalarTower.coe_toAlgHom', IntermediateField.algebraMap_apply] end separableClosure diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean b/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean index 31ff5e240b2d7..89d21876c78f7 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean @@ -147,7 +147,7 @@ theorem isLocalStructomorphOn_contDiffGroupoid_iff_aux {f : PartialHomeomorph M have he'' : ContMDiffOn I I ⊤ e _ := contMDiffOn_of_mem_contDiffGroupoid he have hc : ContMDiffOn I I ⊤ c _ := contMDiffOn_chart refine (hc'.comp' (he''.comp' hc)).mono ?_ - dsimp [s] + dsimp [s, c, c'] mfld_set_tac have H₂ : EqOn f (c'.symm ∘ e ∘ c) s := by intro y hy @@ -185,19 +185,19 @@ theorem isLocalStructomorphOn_contDiffGroupoid_iff (f : PartialHomeomorph M M') e ∈ contDiffGroupoid ∞ I ∧ e.source ⊆ (c.symm ≫ₕ f ≫ₕ c').source ∧ EqOn (c' ∘ f ∘ c.symm) e e.source ∧ c x ∈ e.source := by - have h1 : c' = chartAt H (f x) := by simp only [f.right_inv hX] + have h1 : c' = chartAt H (f x) := by simp only [x, c', f.right_inv hX] have h2 : c' ∘ f ∘ c.symm = ⇑(c.symm ≫ₕ f ≫ₕ c') := rfl have hcx : c x ∈ c.symm ⁻¹' f.source := by simp only [c, hx, mfld_simps] rw [h2] rw [← h1, h2, PartialHomeomorph.isLocalStructomorphWithinAt_iff'] at hxf · exact hxf hcx - · mfld_set_tac + · dsimp [x, c]; mfld_set_tac · apply Or.inl simp only [c, hx, h1, mfld_simps] have h2X : c' X = e (c (f.symm X)) := by rw [← hef hex] dsimp only [Function.comp_def] - have hfX : f.symm X ∈ c.source := by simp only [c, hX, mfld_simps] + have hfX : f.symm X ∈ c.source := by simp only [c, x, hX, mfld_simps] rw [c.left_inv hfX, f.right_inv hX] have h3e : EqOn (c ∘ f.symm ∘ c'.symm) e.symm (c'.symm ⁻¹' f.target ∩ e.target) := by have h1 : EqOn (c.symm ≫ₕ f ≫ₕ c').symm e.symm (e.target ∩ e.target) := by @@ -230,28 +230,30 @@ theorem isLocalStructomorphOn_contDiffGroupoid_iff (f : PartialHomeomorph M M') simp only [mfld_simps] at hy have H : ContMDiffWithinAt I I ⊤ f (f ≫ₕ c').source ((extChartAt I x).symm y) := by refine (h₁ ((extChartAt I x).symm y) ?_).mono ?_ - · simp only [hy, mfld_simps] + · simp only [c, hy, mfld_simps] · mfld_set_tac - have hy' : (extChartAt I x).symm y ∈ c.source := by simp only [hy, mfld_simps] - have hy'' : f ((extChartAt I x).symm y) ∈ c'.source := by simp only [hy, mfld_simps] + have hy' : (extChartAt I x).symm y ∈ c.source := by simp only [c, hy, mfld_simps] + have hy'' : f ((extChartAt I x).symm y) ∈ c'.source := by + simp only [c, hy, mfld_simps] rw [contMDiffWithinAt_iff_of_mem_source hy' hy''] at H convert H.2.mono _ - · simp only [hy, mfld_simps] - · mfld_set_tac + · simp only [c, hy, mfld_simps] + · dsimp [c, c']; mfld_set_tac · -- smoothness of the candidate local structomorphism in the reverse direction intro y hy simp only [mfld_simps] at hy have H : ContMDiffWithinAt I I ⊤ f.symm (f.symm ≫ₕ c).source ((extChartAt I (f x)).symm y) := by refine (h₂ ((extChartAt I (f x)).symm y) ?_).mono ?_ - · simp only [hy, mfld_simps] + · simp only [c', hy, mfld_simps] · mfld_set_tac - have hy' : (extChartAt I (f x)).symm y ∈ c'.source := by simp only [hy, mfld_simps] - have hy'' : f.symm ((extChartAt I (f x)).symm y) ∈ c.source := by simp only [hy, mfld_simps] + have hy' : (extChartAt I (f x)).symm y ∈ c'.source := by simp only [c', hy, mfld_simps] + have hy'' : f.symm ((extChartAt I (f x)).symm y) ∈ c.source := by + simp only [c', hy, mfld_simps] rw [contMDiffWithinAt_iff_of_mem_source hy' hy''] at H convert H.2.mono _ - · simp only [hy, mfld_simps] - · mfld_set_tac + · simp only [c', hy, mfld_simps] + · dsimp [c, c']; mfld_set_tac -- now check the candidate local structomorphism agrees with `f` where it is supposed to · simp only [mfld_simps]; apply eqOn_refl · simp only [c, c', hx', mfld_simps] diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean b/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean index bbe1eedcdb768..af107053b4705 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean @@ -62,19 +62,20 @@ theorem ContMDiffWithinAt.comp {t : Set M'} {g : M' → M''} (x : M) filter_upwards [hf.1.tendsto (extChartAt_source_mem_nhds (I := I') (f x)), inter_mem_nhdsWithin s (extChartAt_source_mem_nhds (I := I) x)] rintro x' (hfx' : f x' ∈ e'.source) ⟨hx's, hx'⟩ - simp only [e.map_source hx', true_and, e.left_inv hx', st hx's, *] + simp only [e, e.map_source hx', true_and, e.left_inv hx', st hx's, *] refine ((hg.2.comp _ (hf.2.mono inter_subset_right) ((mapsTo_preimage _ _).mono_left inter_subset_left)).mono_of_mem_nhdsWithin (inter_mem ?_ self_mem_nhdsWithin)).congr_of_eventuallyEq ?_ ?_ · filter_upwards [A] rintro x' ⟨ht, hfx'⟩ - simp only [*, mem_preimage, writtenInExtChartAt, (· ∘ ·), mem_inter_iff, e'.left_inv, + simp only [*, e, e',mem_preimage, writtenInExtChartAt, (· ∘ ·), mem_inter_iff, e'.left_inv, true_and] exact mem_range_self _ · filter_upwards [A] rintro x' ⟨-, hfx'⟩ - simp only [*, (· ∘ ·), writtenInExtChartAt, e'.left_inv] - · simp only [e, e', writtenInExtChartAt, (· ∘ ·), mem_extChartAt_source, e.left_inv, e'.left_inv] + simp only [*, e, e', (· ∘ ·), writtenInExtChartAt, e'.left_inv] + · simp only [e, e', writtenInExtChartAt, (· ∘ ·), mem_extChartAt_source, + e.left_inv, e'.left_inv] /-- See note [comp_of_eq lemmas] -/ theorem ContMDiffWithinAt.comp_of_eq {t : Set M'} {g : M' → M''} {x : M} {y : M'} diff --git a/Mathlib/LinearAlgebra/Matrix/Charpoly/Coeff.lean b/Mathlib/LinearAlgebra/Matrix/Charpoly/Coeff.lean index e5d7288e96387..5fb0fce3c9ebc 100644 --- a/Mathlib/LinearAlgebra/Matrix/Charpoly/Coeff.lean +++ b/Mathlib/LinearAlgebra/Matrix/Charpoly/Coeff.lean @@ -374,7 +374,7 @@ lemma isNilpotent_charpoly_sub_pow_of_isNilpotent (hM : IsNilpotent M) : have aux : (M.charpoly - X ^ (Fintype.card n)).natDegree ≤ M.charpoly.natDegree := le_trans (natDegree_sub_le _ _) (by simp) rw [← isNilpotent_reflect_iff aux, reflect_sub, ← reverse, M.reverse_charpoly] - simpa [hp] + simpa [p, hp] end reverse diff --git a/Mathlib/LinearAlgebra/SesquilinearForm.lean b/Mathlib/LinearAlgebra/SesquilinearForm.lean index 19fb599ac6529..e4eb407fc013c 100644 --- a/Mathlib/LinearAlgebra/SesquilinearForm.lean +++ b/Mathlib/LinearAlgebra/SesquilinearForm.lean @@ -731,7 +731,7 @@ lemma disjoint_ker_of_nondegenerate_restrict {B : M →ₗ[R] M →ₗ[R] M₁} simp_rw [Subtype.forall, domRestrict₁₂_apply] intro y hy rw [mem_ker] at hx' - simp [hx'] + simp [x', hx'] lemma IsSymm.nondegenerate_restrict_of_isCompl_ker {B : M →ₗ[R] M →ₗ[R] R} (hB : B.IsSymm) {W : Submodule R M} (hW : IsCompl W (LinearMap.ker B)) : diff --git a/Mathlib/MeasureTheory/Constructions/Pi.lean b/Mathlib/MeasureTheory/Constructions/Pi.lean index f7a53cd9ee56e..fb04b17543a7c 100644 --- a/Mathlib/MeasureTheory/Constructions/Pi.lean +++ b/Mathlib/MeasureTheory/Constructions/Pi.lean @@ -500,7 +500,7 @@ lemma pi_map_piOptionEquivProd {β : Option ι → Type*} [∀ i, MeasurableSpac simp only [mem_preimage, Set.mem_pi, mem_univ, forall_true_left, mem_prod] refine ⟨by tauto, fun _ i ↦ ?_⟩ rcases i <;> tauto - simp only [me.map_apply, univ_option, le_eq_subset, Finset.prod_insertNone, this, + simp only [e_meas, me.map_apply, univ_option, le_eq_subset, Finset.prod_insertNone, this, prod_prod, pi_pi, mul_comm] section Intervals diff --git a/Mathlib/MeasureTheory/Covering/Differentiation.lean b/Mathlib/MeasureTheory/Covering/Differentiation.lean index eb17b94c41db8..672bfe1a4f090 100644 --- a/Mathlib/MeasureTheory/Covering/Differentiation.lean +++ b/Mathlib/MeasureTheory/Covering/Differentiation.lean @@ -536,7 +536,7 @@ theorem withDensity_le_mul {s : Set α} (hs : MeasurableSet s) {t : ℝ≥0} (ht have A : ν (s ∩ f ⁻¹' {0}) ≤ ((t : ℝ≥0∞) ^ 2 • ρ :) (s ∩ f ⁻¹' {0}) := by apply le_trans _ (zero_le _) have M : MeasurableSet (s ∩ f ⁻¹' {0}) := hs.inter (f_meas (measurableSet_singleton _)) - simp only [ν, nonpos_iff_eq_zero, M, withDensity_apply, lintegral_eq_zero_iff f_meas] + simp only [f, ν, nonpos_iff_eq_zero, M, withDensity_apply, lintegral_eq_zero_iff f_meas] apply (ae_restrict_iff' M).2 exact Eventually.of_forall fun x hx => hx.2 have B : ν (s ∩ f ⁻¹' {∞}) ≤ ((t : ℝ≥0∞) ^ 2 • ρ :) (s ∩ f ⁻¹' {∞}) := by @@ -551,7 +551,7 @@ theorem withDensity_le_mul {s : Set α} (hs : MeasurableSet s) {t : ℝ≥0} (ht intro n let I := Ico ((t : ℝ≥0∞) ^ n) ((t : ℝ≥0∞) ^ (n + 1)) have M : MeasurableSet (s ∩ f ⁻¹' I) := hs.inter (f_meas measurableSet_Ico) - simp only [ν, M, withDensity_apply, coe_nnreal_smul_apply] + simp only [ν, I, M, withDensity_apply, coe_nnreal_smul_apply] calc (∫⁻ x in s ∩ f ⁻¹' I, f x ∂μ) ≤ ∫⁻ _ in s ∩ f ⁻¹' I, (t : ℝ≥0∞) ^ (n + 1) ∂μ := lintegral_mono_ae ((ae_restrict_iff' M).2 (Eventually.of_forall fun x hx => hx.2.2.le)) @@ -614,7 +614,7 @@ theorem le_mul_withDensity {s : Set α} (hs : MeasurableSet s) {t : ℝ≥0} (ht intro n let I := Ico ((t : ℝ≥0∞) ^ n) ((t : ℝ≥0∞) ^ (n + 1)) have M : MeasurableSet (s ∩ f ⁻¹' I) := hs.inter (f_meas measurableSet_Ico) - simp only [ν, M, withDensity_apply, coe_nnreal_smul_apply] + simp only [ν, I, M, withDensity_apply, coe_nnreal_smul_apply] calc ρ (s ∩ f ⁻¹' I) ≤ (t : ℝ≥0∞) ^ (n + 1) * μ (s ∩ f ⁻¹' I) := by rw [← ENNReal.coe_zpow t_ne_zero'] @@ -736,7 +736,7 @@ theorem ae_tendsto_measure_inter_div (s : Set α) : have B : ∀ᵐ x ∂μ.restrict s, t.indicator 1 x = (1 : ℝ≥0∞) := by refine ae_restrict_of_ae_restrict_of_subset (subset_toMeasurable μ s) ?_ filter_upwards [ae_restrict_mem (measurableSet_toMeasurable μ s)] with _ hx - simp only [hx, Pi.one_apply, indicator_of_mem] + simp only [t, hx, Pi.one_apply, indicator_of_mem] filter_upwards [A, B] with x hx h'x rw [h'x] at hx apply hx.congr' _ diff --git a/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean b/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean index 2365cd675affa..4ebbecf29aa5d 100644 --- a/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean +++ b/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean @@ -97,7 +97,7 @@ theorem SimpleFunc.exists_le_lowerSemicontinuous_lintegral_ge (f : α →ₛ ℝ by_cases h : ∫⁻ x, f x ∂μ = ⊤ · refine ⟨fun _ => c, fun x => ?_, lowerSemicontinuous_const, by - simp only [_root_.top_add, le_top, h]⟩ + simp only [f, _root_.top_add, le_top, h]⟩ simp only [SimpleFunc.coe_const, SimpleFunc.const_zero, SimpleFunc.coe_zero, Set.piecewise_eq_indicator, SimpleFunc.coe_piecewise] exact Set.indicator_le_self _ _ _ diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean index fc756040cc012..90b9c3cc0537b 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean @@ -82,7 +82,7 @@ theorem MeasureTheory.measure_lt_one_eq_integral_div_gamma {p : ℝ} (hp : 0 < p dist := fun x y => g (x - y) dist_self := by simp only [_root_.sub_self, h1, forall_const] dist_comm := fun _ _ => by dsimp [dist]; rw [← h2, neg_sub] - dist_triangle := fun x y z => by convert h3 (x - y) (y - z) using 1; abel_nf + dist_triangle := fun x y z => by convert h3 (x - y) (y - z) using 1; simp [F] edist := fun x y => .ofReal (g (x - y)) edist_dist := fun _ _ => rfl eq_of_dist_eq_zero := by convert fun _ _ h => eq_of_sub_eq_zero (h4 h) } @@ -125,7 +125,7 @@ theorem MeasureTheory.measure_le_eq_lt [Nontrivial E] (r : ℝ) : dist := fun x y => g (x - y) dist_self := by simp only [_root_.sub_self, h1, forall_const] dist_comm := fun _ _ => by dsimp [dist]; rw [← h2, neg_sub] - dist_triangle := fun x y z => by convert h3 (x - y) (y - z) using 1; abel_nf + dist_triangle := fun x y z => by convert h3 (x - y) (y - z) using 1; simp [F] edist := fun x y => .ofReal (g (x - y)) edist_dist := fun _ _ => rfl eq_of_dist_eq_zero := by convert fun _ _ h => eq_of_sub_eq_zero (h4 h) } diff --git a/Mathlib/NumberTheory/Bernoulli.lean b/Mathlib/NumberTheory/Bernoulli.lean index 28dc748d180d4..dbd30c72d4719 100644 --- a/Mathlib/NumberTheory/Bernoulli.lean +++ b/Mathlib/NumberTheory/Bernoulli.lean @@ -285,7 +285,7 @@ theorem sum_range_pow (n p : ℕ) : ext q : 1 let f a b := bernoulli a / a ! * coeff ℚ (b + 1) (exp ℚ ^ n) -- key step: use `PowerSeries.coeff_mul` and then rewrite sums - simp only [coeff_mul, coeff_mk, cast_mul, sum_antidiagonal_eq_sum_range_succ f] + simp only [f, coeff_mul, coeff_mk, cast_mul, sum_antidiagonal_eq_sum_range_succ f] apply sum_congr rfl intros m h simp only [f, exp_pow_eq_rescale_exp, rescale, one_div, coeff_mk, RingHom.coe_mk, coeff_exp, diff --git a/Mathlib/NumberTheory/EulerProduct/Basic.lean b/Mathlib/NumberTheory/EulerProduct/Basic.lean index be623e1cac957..590d05ced1a79 100644 --- a/Mathlib/NumberTheory/EulerProduct/Basic.lean +++ b/Mathlib/NumberTheory/EulerProduct/Basic.lean @@ -208,7 +208,7 @@ theorem eulerProduct (hsum : Summable (‖f ·‖)) (hf₀ : f 0 = 0) : have H (n : ℕ) : ∏ i ∈ range n, Set.mulIndicator {p | Nat.Prime p} F i = ∏ p ∈ primesBelow n, ∑' (e : ℕ), f (p ^ e) := prod_mulIndicator_eq_prod_filter (range n) (fun _ ↦ F) (fun _ ↦ {p | Nat.Prime p}) id - simpa only [H] + simpa only [F, H] include hf₁ hmul in /-- The *Euler Product* for multiplicative (on coprime arguments) functions. diff --git a/Mathlib/NumberTheory/WellApproximable.lean b/Mathlib/NumberTheory/WellApproximable.lean index a4baa65fee7dd..f12024dfbf347 100644 --- a/Mathlib/NumberTheory/WellApproximable.lean +++ b/Mathlib/NumberTheory/WellApproximable.lean @@ -307,9 +307,9 @@ theorem addWellApproximable_ae_empty_or_univ (δ : ℕ → ℝ) (hδ : Tendsto rw [hE₁ p] cases hp · cases' hA p with _ h; · contradiction - simp only [h, union_ae_eq_univ_of_ae_eq_univ_left] + simp only [μ, h, union_ae_eq_univ_of_ae_eq_univ_left] · cases' hB p with _ h; · contradiction - simp only [h, union_ae_eq_univ_of_ae_eq_univ_left, + simp only [μ, h, union_ae_eq_univ_of_ae_eq_univ_left, union_ae_eq_univ_of_ae_eq_univ_right] /-- A general version of **Dirichlet's approximation theorem**. diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index ddfb081b4e76b..cf1e32cc662bd 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -435,7 +435,7 @@ theorem iInf_sets_eq {f : ι → Filter α} (h : Directed (· ≥ ·) f) [ne : N exact ⟨c, inter_mem (ha hx) (hb hy)⟩ } have : u = iInf f := eq_iInf_of_mem_iff_exists_mem mem_iUnion -- Porting note: it was just `congr_arg filter.sets this.symm` - (congr_arg Filter.sets this.symm).trans <| by simp only + (congr_arg Filter.sets this.symm).trans <| by simp only [u] theorem mem_iInf_of_directed {f : ι → Filter α} (h : Directed (· ≥ ·) f) [Nonempty ι] (s) : s ∈ iInf f ↔ ∃ i, s ∈ f i := by diff --git a/Mathlib/Order/LiminfLimsup.lean b/Mathlib/Order/LiminfLimsup.lean index 6cdea2ca4d772..1b33f08a0dec9 100644 --- a/Mathlib/Order/LiminfLimsup.lean +++ b/Mathlib/Order/LiminfLimsup.lean @@ -1430,7 +1430,7 @@ theorem HasBasis.liminf_eq_ciSup_ciInf {v : Filter ι} apply Subset.antisymm · apply iUnion_subset (fun j ↦ ?_) by_cases hj : j ∈ m - · have : j = liminf_reparam f s p j := by simp only [liminf_reparam, hj, ite_true] + · have : j = liminf_reparam f s p j := by simp only [m, liminf_reparam, hj, ite_true] conv_lhs => rw [this] apply subset_iUnion _ j · simp only [m, mem_setOf_eq, ← nonempty_iInter_Iic_iff, not_nonempty_iff_eq_empty] at hj @@ -1443,8 +1443,8 @@ theorem HasBasis.liminf_eq_ciSup_ciInf {v : Filter ι} apply (Iic_ciInf _).symm change liminf_reparam f s p j ∈ m by_cases Hj : j ∈ m - · simpa only [liminf_reparam, if_pos Hj] using Hj - · simp only [liminf_reparam, if_neg Hj] + · simpa only [m, liminf_reparam, if_pos Hj] using Hj + · simp only [m, liminf_reparam, if_neg Hj] have Z : ∃ n, (exists_surjective_nat (Subtype p)).choose n ∈ m ∨ ∀ j, j ∉ m := by rcases (exists_surjective_nat (Subtype p)).choose_spec j0 with ⟨n, rfl⟩ exact ⟨n, Or.inl hj0⟩ diff --git a/Mathlib/Order/Partition/Equipartition.lean b/Mathlib/Order/Partition/Equipartition.lean index 648a5421e9ba3..4aec6bb5e03ea 100644 --- a/Mathlib/Order/Partition/Equipartition.lean +++ b/Mathlib/Order/Partition/Equipartition.lean @@ -147,7 +147,7 @@ theorem IsEquipartition.exists_partPreservingEquiv (hP : P.IsEquipartition) : exact Sigma.ext e.2 <| (Fin.heq_ext_iff (by rw [e.2])).mpr e.1 use Equiv.ofBijective _ bij intro a b - simp_rw [Equiv.ofBijective_apply, z, hf a b, Nat.mul_add_mod, + simp_rw [z', z, Equiv.ofBijective_apply, hf a b, Nat.mul_add_mod, Nat.mod_eq_of_lt (gl a), Nat.mod_eq_of_lt (gl b), Fin.val_eq_val, g.apply_eq_iff_eq] /-! ### Discrete and indiscrete finpartitions -/ diff --git a/Mathlib/Probability/Integration.lean b/Mathlib/Probability/Integration.lean index 3d42e69819c63..e289b19f7beac 100644 --- a/Mathlib/Probability/Integration.lean +++ b/Mathlib/Probability/Integration.lean @@ -142,6 +142,7 @@ theorem IndepFun.integrable_mul {β : Type*} [MeasurableSpace β] {X Y : Ω → have hmul : ∫⁻ a, nX a * nY a ∂μ = (∫⁻ a, nX a ∂μ) * ∫⁻ a, nY a ∂μ := lintegral_mul_eq_lintegral_mul_lintegral_of_indepFun' hnX hnY hXY'' refine ⟨hX.1.mul hY.1, ?_⟩ + simp only [nX, nY] at hmul simp_rw [HasFiniteIntegral, Pi.mul_apply, nnnorm_mul, ENNReal.coe_mul, hmul] exact ENNReal.mul_lt_top hX.2 hY.2 diff --git a/Mathlib/Probability/Process/Stopping.lean b/Mathlib/Probability/Process/Stopping.lean index 40bf2069437c8..eb62d512ac6fa 100644 --- a/Mathlib/Probability/Process/Stopping.lean +++ b/Mathlib/Probability/Process/Stopping.lean @@ -752,7 +752,7 @@ theorem progMeasurable_min_stopping_time [MetrizableSpace ι] (hτ : IsStoppingT suffices h_min_eq_left : (fun x : sc => min (↑(x : Set.Iic i × Ω).fst) (τ (x : Set.Iic i × Ω).snd)) = fun x : sc => ↑(x : Set.Iic i × Ω).fst by - simp (config := { unfoldPartialApp := true }) only [Set.restrict, h_min_eq_left] + simp +unfoldPartialApp only [sc, Set.restrict, h_min_eq_left] exact h_meas_fst _ ext1 ω rw [min_eq_left] diff --git a/Mathlib/Probability/StrongLaw.lean b/Mathlib/Probability/StrongLaw.lean index 9094e67f68a42..8c69211cebfef 100644 --- a/Mathlib/Probability/StrongLaw.lean +++ b/Mathlib/Probability/StrongLaw.lean @@ -435,7 +435,7 @@ theorem strong_law_aux1 {c : ℝ} (c_one : 1 < c) {ε : ℝ} (εpos : 0 < ε) : simp only [Y, Nat.cast_zero, truncation_zero, variance_zero, mul_zero, le_rfl] apply mul_le_mul_of_nonneg_right _ (variance_nonneg _ _) convert sum_div_nat_floor_pow_sq_le_div_sq N (Nat.cast_pos.2 hj) c_one using 2 - · simp only [Nat.cast_lt] + · simp only [u, Nat.cast_lt] · simp only [Y, S, u, C, one_div] _ = c ^ 5 * (c - 1)⁻¹ ^ 3 * ∑ j ∈ range (u (N - 1)), ((j : ℝ) ^ 2)⁻¹ * Var[Y j] := by simp_rw [mul_sum, div_eq_mul_inv, mul_assoc] @@ -634,7 +634,7 @@ theorem strong_law_ae_real {Ω : Type*} {m : MeasurableSpace Ω} {μ : Measure convert hωpos.sub hωneg using 2 · simp only [pos, neg, ← sub_div, ← sum_sub_distrib, max_zero_sub_max_neg_zero_eq_self, Function.comp_apply] - · simp only [← integral_sub hint.pos_part hint.neg_part, + · simp only [pos, neg, ← integral_sub hint.pos_part hint.neg_part, max_zero_sub_max_neg_zero_eq_self, Function.comp_apply, mΩ] end StrongLawAeReal diff --git a/Mathlib/RingTheory/FinitePresentation.lean b/Mathlib/RingTheory/FinitePresentation.lean index 64e20502c1f9e..777df2c59cb87 100644 --- a/Mathlib/RingTheory/FinitePresentation.lean +++ b/Mathlib/RingTheory/FinitePresentation.lean @@ -364,7 +364,7 @@ theorem ker_fg_of_mvPolynomial {n : ℕ} (f : MvPolynomial (Fin n) R →ₐ[R] A · exact add_mem (Ideal.mul_mem_right _ _ hy₁) (Ideal.mul_mem_left _ _ hy₂) obtain ⟨_, ⟨x, rfl⟩, y, hy, rfl⟩ := AddSubmonoid.mem_sup.mp this refine add_mem ?_ hy - simp only [RingHom.mem_ker, AlgHom.toRingHom_eq_coe, AlgHom.coe_toRingHom, map_add, + simp only [RXn, RingHom.mem_ker, AlgHom.toRingHom_eq_coe, AlgHom.coe_toRingHom, map_add, show f y = 0 from leI hy, add_zero, hh'] at hx suffices Ideal.span (s : Set RXm) ≤ (Ideal.span s').comap aeval_h by apply this diff --git a/Mathlib/RingTheory/Flat/Basic.lean b/Mathlib/RingTheory/Flat/Basic.lean index 18985920c6140..964c6fdd3e82a 100644 --- a/Mathlib/RingTheory/Flat/Basic.lean +++ b/Mathlib/RingTheory/Flat/Basic.lean @@ -205,8 +205,8 @@ instance directSum (ι : Type v) (M : ι → Type w) [(i : ι) → AddCommGroup have h₂ := F i rw [iff_rTensor_injective] at h₂ have h₃ := h₂ hI - simp only [coe_comp, LinearEquiv.coe_coe, Function.comp_apply, EmbeddingLike.map_eq_zero_iff, - h₃, LinearMap.map_eq_zero_iff] at f + simp only [φ, τ, coe_comp, LinearEquiv.coe_coe, Function.comp_apply, + EmbeddingLike.map_eq_zero_iff, h₃, LinearMap.map_eq_zero_iff] at f simp [f] open scoped Classical in diff --git a/Mathlib/RingTheory/Jacobson/Ring.lean b/Mathlib/RingTheory/Jacobson/Ring.lean index 958012570ae25..c2b35f330796f 100644 --- a/Mathlib/RingTheory/Jacobson/Ring.lean +++ b/Mathlib/RingTheory/Jacobson/Ring.lean @@ -306,7 +306,7 @@ theorem isIntegral_isLocalization_polynomial_quotient obtain ⟨q'', hq''⟩ := isUnit_iff_exists_inv'.1 (IsLocalization.map_units Rₘ (⟨q', hq'⟩ : M)) refine (hp.symm ▸ this).of_mul_unit φ' p (algebraMap (R[X] ⧸ P) Sₘ (φ q')) q'' ?_ rw [← φ'.map_one, ← congr_arg φ' hq'', φ'.map_mul, ← φ'.comp_apply] - simp only [IsLocalization.map_comp _] + simp only [φ', IsLocalization.map_comp _] rw [RingHom.comp_apply] dsimp at hp refine @IsIntegral.of_mem_closure'' Rₘ _ Sₘ _ φ' @@ -320,7 +320,7 @@ theorem isIntegral_isLocalization_polynomial_quotient (pX.map (Ideal.Quotient.mk P')) ?_ M ?_ · rwa [eval₂_map, hφ', ← hom_eval₂, Quotient.eq_zero_iff_mem, eval₂_C_X] · use 1 - simp only [pow_one] + simp only [P', pow_one] · rw [Set.mem_setOf_eq, degree_le_zero_iff] at hy -- Porting note: was `refine' hy.symm ▸` -- `⟨X - C (algebraMap _ _ ((Quotient.mk P') (p.coeff 0))), monic_X_sub_C _, _⟩` @@ -330,7 +330,7 @@ theorem isIntegral_isLocalization_polynomial_quotient · apply monic_X_sub_C · simp only [eval₂_sub, eval₂_X, eval₂_C] rw [sub_eq_zero, ← φ'.comp_apply] - simp only [IsLocalization.map_comp _] + simp only [φ', IsLocalization.map_comp _] rfl · obtain ⟨p, rfl⟩ := Ideal.Quotient.mk_surjective p' rw [← RingHom.comp_apply] diff --git a/Mathlib/RingTheory/Localization/Ideal.lean b/Mathlib/RingTheory/Localization/Ideal.lean index 5e156755a12ed..fc1af5344f892 100644 --- a/Mathlib/RingTheory/Localization/Ideal.lean +++ b/Mathlib/RingTheory/Localization/Ideal.lean @@ -67,7 +67,7 @@ theorem mem_map_algebraMap_iff {I : Ideal R} {z} : z ∈ Ideal.map (algebraMap R obtain ⟨y, hy⟩ := hz let Z : { x // x ∈ I } := ⟨y, hy.left⟩ use ⟨Z, 1⟩ - simp [hy.right] + simp [Z, hy.right] · rintro ⟨⟨a, s⟩, h⟩ rw [← Ideal.unit_mul_mem_iff_mem _ (map_units S s), mul_comm] exact h.symm ▸ Ideal.mem_map_of_mem _ a.2 diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean index 2964fdd687505..eadb802314521 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean @@ -241,8 +241,8 @@ theorem sum_antidiagonal_card_esymm_psum_eq_zero : suffices (-1 : MvPolynomial σ R) ^ (k + 1) * ∑ a ∈ antidiagonal k, (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd = 0 by simpa using this - 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)] + simp [k, ← 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)] /-- 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. -/ diff --git a/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean b/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean index 6365d391b724f..ca5b849f7101c 100644 --- a/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean +++ b/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean @@ -145,7 +145,7 @@ theorem dvd_coeff_zero_of_aeval_eq_prime_smul_of_minpoly_isEisensteinAt {B : Pow choose! f hf using hei.isWeaklyEisensteinAt.exists_mem_adjoin_mul_eq_pow_natDegree_le (minpoly.aeval R B.gen) (minpoly.monic hBint) - simp only [(minpoly.monic hBint).natDegree_map, deg_R_P] at hf + simp only [P, (minpoly.monic hBint).natDegree_map, deg_R_P] at hf -- The Eisenstein condition shows that `p` divides `Q.coeff 0` -- if `p^n.succ` divides the following multiple of `Q.coeff 0^n.succ`: diff --git a/Mathlib/RingTheory/Polynomial/Nilpotent.lean b/Mathlib/RingTheory/Polynomial/Nilpotent.lean index a921d48c10bc8..8076374cc6e73 100644 --- a/Mathlib/RingTheory/Polynomial/Nilpotent.lean +++ b/Mathlib/RingTheory/Polynomial/Nilpotent.lean @@ -119,10 +119,10 @@ theorem isUnit_of_coeff_isUnit_isNilpotent (hunit : IsUnit (P.coeff 0)) have hdeg₂ := lt_of_le_of_lt P.eraseLead_natDegree_le (Nat.sub_lt (Nat.pos_of_ne_zero hdeg) zero_lt_one) refine hind P₁.natDegree ?_ ?_ (fun i hi => ?_) rfl - · simp_rw [← h, hdeg₂] - · simp_rw [eraseLead_coeff_of_ne _ (Ne.symm hdeg), hunit] + · simp_rw [P₁, ← h, hdeg₂] + · simp_rw [P₁, eraseLead_coeff_of_ne _ (Ne.symm hdeg), hunit] · by_cases H : i ≤ P₁.natDegree - · simp_rw [eraseLead_coeff_of_ne _ (ne_of_lt (lt_of_le_of_lt H hdeg₂)), hnil i hi] + · simp_rw [P₁, eraseLead_coeff_of_ne _ (ne_of_lt (lt_of_le_of_lt H hdeg₂)), hnil i hi] · simp_rw [coeff_eq_zero_of_natDegree_lt (lt_of_not_ge H), IsNilpotent.zero] /-- Let `P` be a polynomial over `R`. If `P` is a unit, then all its coefficients are nilpotent, diff --git a/Mathlib/RingTheory/RingHom/Locally.lean b/Mathlib/RingTheory/RingHom/Locally.lean index 6dd45ad1c715d..e3370472307ba 100644 --- a/Mathlib/RingTheory/RingHom/Locally.lean +++ b/Mathlib/RingTheory/RingHom/Locally.lean @@ -218,7 +218,7 @@ lemma locally_stableUnderComposition (hPi : RespectsIso P) (hPl : LocalizationPr change _ = Localization.awayMap g' a.val (algebraMap S _ (f x)) simp only [Localization.awayMap, IsLocalization.Away.map, IsLocalization.map_eq] rfl - simp only [this] + simp only [this, a'] apply hPc _ _ (hsf a.val a.property) apply @hPl _ _ _ _ g' _ _ _ _ _ _ _ _ ?_ (hsg b.val b.property) exact IsLocalization.Away.instMapRingHomPowersOfCoe (Localization.Away (g' a.val)) a.val @@ -323,10 +323,10 @@ lemma locally_localizationAwayPreserves (hPl : LocalizationAwayPreserves P) : haveI (a : s) : IsLocalization (Algebra.algebraMapSubmonoid (Localization.Away a.val) (Submonoid.map f (Submonoid.powers r))) (Sₐ a) := by convert inferInstanceAs (IsLocalization.Away (rₐ a) (Sₐ a)) - simp [Algebra.algebraMapSubmonoid] + simp [rₐ, Sₐ, Algebra.algebraMapSubmonoid] have H (a : s) : Submonoid.powers (f r) ≤ (Submonoid.powers (rₐ a)).comap (algebraMap S (Localization.Away a.val)) := by - simp [Submonoid.powers_le] + simp [rₐ, Sₐ, Submonoid.powers_le] letI (a : s) : Algebra S' (Sₐ a) := (IsLocalization.map (Sₐ a) (algebraMap S (Localization.Away a.val)) (H a)).toAlgebra haveI (a : s) : IsScalarTower S S' (Sₐ a) := diff --git a/Mathlib/RingTheory/Smooth/Kaehler.lean b/Mathlib/RingTheory/Smooth/Kaehler.lean index dc6e1c760701e..780d71c547b67 100644 --- a/Mathlib/RingTheory/Smooth/Kaehler.lean +++ b/Mathlib/RingTheory/Smooth/Kaehler.lean @@ -347,7 +347,7 @@ def retractionKerCotangentToTensorEquivSection : fun ⟨l, hl⟩ ↦ ⟨e₁.symm.toLinearMap ∘ₗ l.restrictScalars P ∘ₗ e₂.symm.toLinearMap, ?_⟩, ?_, ?_⟩ · rintro x y obtain ⟨x, rfl⟩ := Ideal.Quotient.mk_surjective x - simp only [← Ideal.Quotient.algebraMap_eq, IsScalarTower.algebraMap_smul] + simp only [P', ← Ideal.Quotient.algebraMap_eq, IsScalarTower.algebraMap_smul] exact (e₁.toLinearMap ∘ₗ l ∘ₗ e₂.toLinearMap).map_smul x y · ext1 x rw [H] at hl diff --git a/Mathlib/RingTheory/Trace/Quotient.lean b/Mathlib/RingTheory/Trace/Quotient.lean index dd0fb701f5e47..2ab1cd6fe80de 100644 --- a/Mathlib/RingTheory/Trace/Quotient.lean +++ b/Mathlib/RingTheory/Trace/Quotient.lean @@ -207,8 +207,8 @@ lemma Algebra.trace_quotient_eq_of_isDedekindDomain (x) [IsDedekindDomain R] [Is RingHom.ker_eq_bot_iff_eq_zero] intro x hx obtain ⟨x, s, rfl⟩ := IsLocalization.mk'_surjective p.primeCompl x - simp only [RingHom.algebraMap_toAlgebra, IsLocalization.map_mk', IsLocalization.mk'_eq_zero_iff, - mul_eq_zero, Subtype.exists, exists_prop] at hx ⊢ + simp only [Sₚ, RingHom.algebraMap_toAlgebra, IsLocalization.map_mk', + IsLocalization.mk'_eq_zero_iff, mul_eq_zero, Subtype.exists, exists_prop] at hx ⊢ obtain ⟨_, ⟨a, ha, rfl⟩, H⟩ := hx simp only [(injective_iff_map_eq_zero' _).mp (NoZeroSMulDivisors.algebraMap_injective R S)] at H refine ⟨a, ha, H⟩ diff --git a/Mathlib/RingTheory/Unramified/Field.lean b/Mathlib/RingTheory/Unramified/Field.lean index 37f53dffd595d..1c82e3e8fc22d 100644 --- a/Mathlib/RingTheory/Unramified/Field.lean +++ b/Mathlib/RingTheory/Unramified/Field.lean @@ -179,8 +179,10 @@ theorem range_eq_top_of_isPurelyInseparable 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 + let a : S := ⟨1, h.subset_extend _ (by simp)⟩ + have ha : Basis.extend h a = 1 := by simp [a] + let b : S := ⟨x, h.subset_extend _ (by simp)⟩ + have hb : Basis.extend h b = x := by simp [b] by_cases e : a = b · obtain rfl : 1 = x := congr_arg Subtype.val e exact ⟨1, map_one _⟩ diff --git a/Mathlib/RingTheory/Unramified/Finite.lean b/Mathlib/RingTheory/Unramified/Finite.lean index 59c67331b55c7..7abca42bd9322 100644 --- a/Mathlib/RingTheory/Unramified/Finite.lean +++ b/Mathlib/RingTheory/Unramified/Finite.lean @@ -109,13 +109,14 @@ lemma finite_of_free_aux (I) [DecidableEq I] (b : Basis I R S) · intros; simp only [add_smul] have h₂ : ∀ (x : S), ((b.repr x).support.sum fun a ↦ b.repr x a • b a) = x := by simpa only [Finsupp.linearCombination_apply, Finsupp.sum] using b.linearCombination_repr + simp only [a] at h₁ simp_rw [map_finsupp_sum, map_smul, h₁, Finsupp.sum, Finset.sum_comm (t := f.support), TensorProduct.smul_tmul', ← TensorProduct.sum_tmul, ← Finset.smul_sum, h₂] apply Finset.sum_congr rfl intros i hi apply Finset.sum_subset_zero_on_sdiff · exact Finset.subset_biUnion_of_mem (fun i ↦ (a i).support) hi - · simp only [Finset.mem_sdiff, Finset.mem_biUnion, Finsupp.mem_support_iff, ne_eq, not_not, + · simp only [a, Finset.mem_sdiff, Finset.mem_biUnion, Finsupp.mem_support_iff, ne_eq, not_not, and_imp, forall_exists_index] simp (config := {contextual := true}) · exact fun _ _ ↦ rfl diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Real.lean b/Mathlib/Topology/Algebra/InfiniteSum/Real.lean index 76bd90df592fd..cfd31f628c2ee 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Real.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Real.lean @@ -57,7 +57,7 @@ section summable theorem not_summable_iff_tendsto_nat_atTop_of_nonneg {f : ℕ → ℝ} (hf : ∀ n, 0 ≤ f n) : ¬Summable f ↔ Tendsto (fun n : ℕ => ∑ i ∈ Finset.range n, f i) atTop atTop := by lift f to ℕ → ℝ≥0 using hf - exact mod_cast NNReal.not_summable_iff_tendsto_nat_atTop + simpa using mod_cast NNReal.not_summable_iff_tendsto_nat_atTop theorem summable_iff_not_tendsto_nat_atTop_of_nonneg {f : ℕ → ℝ} (hf : ∀ n, 0 ≤ f n) : Summable f ↔ ¬Tendsto (fun n : ℕ => ∑ i ∈ Finset.range n, f i) atTop atTop := by @@ -66,7 +66,7 @@ theorem summable_iff_not_tendsto_nat_atTop_of_nonneg {f : ℕ → ℝ} (hf : ∀ theorem summable_sigma_of_nonneg {α} {β : α → Type*} {f : (Σ x, β x) → ℝ} (hf : ∀ x, 0 ≤ f x) : Summable f ↔ (∀ x, Summable fun y => f ⟨x, y⟩) ∧ Summable fun x => ∑' y, f ⟨x, y⟩ := by lift f to (Σx, β x) → ℝ≥0 using hf - exact mod_cast NNReal.summable_sigma + simpa using mod_cast NNReal.summable_sigma lemma summable_partition {α β : Type*} {f : β → ℝ} (hf : 0 ≤ f) {s : α → Set β} (hs : ∀ i, ∃! j, i ∈ s j) : Summable f ↔ diff --git a/Mathlib/Topology/Category/Stonean/Basic.lean b/Mathlib/Topology/Category/Stonean/Basic.lean index 58018bc5c4a57..9df5667c0cf2e 100644 --- a/Mathlib/Topology/Category/Stonean/Basic.lean +++ b/Mathlib/Topology/Category/Stonean/Basic.lean @@ -133,7 +133,7 @@ lemma epi_iff_surjective {X Y : Stonean} (f : X ⟶ Y) : have hC : IsClosed C := (isCompact_range f.continuous).isClosed let U := Cᶜ have hUy : U ∈ 𝓝 y := by - simp only [C, Set.mem_range, hy, exists_false, not_false_eq_true, hC.compl_mem_nhds] + simp only [U, C, Set.mem_range, hy, exists_false, not_false_eq_true, hC.compl_mem_nhds] obtain ⟨V, hV, hyV, hVU⟩ := isTopologicalBasis_isClopen.mem_nhds_iff.mp hUy classical let g : Y ⟶ mkFinite (ULift (Fin 2)) := diff --git a/Mathlib/Topology/MetricSpace/GromovHausdorff.lean b/Mathlib/Topology/MetricSpace/GromovHausdorff.lean index 70c260708c6a3..a19d3d5982139 100644 --- a/Mathlib/Topology/MetricSpace/GromovHausdorff.lean +++ b/Mathlib/Topology/MetricSpace/GromovHausdorff.lean @@ -698,20 +698,20 @@ instance : SecondCountableTopology GHSpace := by let i : ℕ := E p x have hip : i < N p := ((E p) x).2 have hiq : i < N q := by rwa [Npq] at hip - have i' : i = (E q) (Ψ x) := by simp only [Ψ, Equiv.apply_symm_apply, Fin.coe_cast] + have i' : i = (E q) (Ψ x) := by simp only [i, Ψ, Equiv.apply_symm_apply, Fin.coe_cast] -- introduce `j`, that codes both `y` and `Φ y` in `Fin (N p) = Fin (N q)` let j : ℕ := E p y have hjp : j < N p := ((E p) y).2 have hjq : j < N q := by rwa [Npq] at hjp have j' : j = ((E q) (Ψ y)).1 := by - simp only [Ψ, Equiv.apply_symm_apply, Fin.coe_cast] + simp only [j, Ψ, Equiv.apply_symm_apply, Fin.coe_cast] -- Express `dist x y` in terms of `F p` have : (F p).2 ((E p) x) ((E p) y) = ⌊ε⁻¹ * dist x y⌋ := by - simp only [(E p).symm_apply_apply] + simp only [F, (E p).symm_apply_apply] have Ap : (F p).2 ⟨i, hip⟩ ⟨j, hjp⟩ = ⌊ε⁻¹ * dist x y⌋ := by rw [← this] -- Express `dist (Φ x) (Φ y)` in terms of `F q` have : (F q).2 ((E q) (Ψ x)) ((E q) (Ψ y)) = ⌊ε⁻¹ * dist (Ψ x) (Ψ y)⌋ := by - simp only [(E q).symm_apply_apply] + simp only [F, (E q).symm_apply_apply] have Aq : (F q).2 ⟨i, hiq⟩ ⟨j, hjq⟩ = ⌊ε⁻¹ * dist (Ψ x) (Ψ y)⌋ := by rw [← this] -- Porting note: `congr` fails to make progress @@ -850,18 +850,18 @@ theorem totallyBounded {t : Set GHSpace} {C : ℝ} {u : ℕ → ℝ} {K : ℕ let i : ℕ := E p x have hip : i < N p := ((E p) x).2 have hiq : i < N q := by rwa [Npq] at hip - have i' : i = (E q) (Ψ x) := by simp only [Ψ, Equiv.apply_symm_apply, Fin.coe_cast] + have i' : i = (E q) (Ψ x) := by simp only [i, Ψ, Equiv.apply_symm_apply, Fin.coe_cast] -- introduce `j`, that codes both `y` and `Φ y` in `Fin (N p) = Fin (N q)` let j : ℕ := E p y have hjp : j < N p := ((E p) y).2 have hjq : j < N q := by rwa [Npq] at hjp - have j' : j = (E q) (Ψ y) := by simp only [Ψ, Equiv.apply_symm_apply, Fin.coe_cast] + have j' : j = (E q) (Ψ y) := by simp only [j, Ψ, Equiv.apply_symm_apply, Fin.coe_cast] -- Express `dist x y` in terms of `F p` have Ap : ((F p).2 ⟨i, hip⟩ ⟨j, hjp⟩).1 = ⌊ε⁻¹ * dist x y⌋₊ := calc ((F p).2 ⟨i, hip⟩ ⟨j, hjp⟩).1 = ((F p).2 ((E p) x) ((E p) y)).1 := by congr - _ = min M ⌊ε⁻¹ * dist x y⌋₊ := by simp only [(E p).symm_apply_apply] + _ = min M ⌊ε⁻¹ * dist x y⌋₊ := by simp only [F, (E p).symm_apply_apply] _ = ⌊ε⁻¹ * dist x y⌋₊ := by refine min_eq_right (Nat.floor_mono ?_) refine mul_le_mul_of_nonneg_left (le_trans ?_ (le_max_left _ _)) (inv_pos.2 εpos).le @@ -874,7 +874,7 @@ theorem totallyBounded {t : Set GHSpace} {C : ℝ} {u : ℕ → ℝ} {K : ℕ ((F q).2 ⟨i, hiq⟩ ⟨j, hjq⟩).1 = ((F q).2 ((E q) (Ψ x)) ((E q) (Ψ y))).1 := by -- Porting note: `congr` drops `Fin.val` but fails to make further progress exact congr_arg₂ (Fin.val <| (F q).2 · ·) (Fin.ext i') (Fin.ext j') - _ = min M ⌊ε⁻¹ * dist (Ψ x) (Ψ y)⌋₊ := by simp only [(E q).symm_apply_apply] + _ = min M ⌊ε⁻¹ * dist (Ψ x) (Ψ y)⌋₊ := by simp only [F, (E q).symm_apply_apply] _ = ⌊ε⁻¹ * dist (Ψ x) (Ψ y)⌋₊ := by refine min_eq_right (Nat.floor_mono ?_) refine mul_le_mul_of_nonneg_left (le_trans ?_ (le_max_left _ _)) (inv_pos.2 εpos).le diff --git a/Mathlib/Topology/MetricSpace/Polish.lean b/Mathlib/Topology/MetricSpace/Polish.lean index 93c053adc16d7..70ae87a47d620 100644 --- a/Mathlib/Topology/MetricSpace/Polish.lean +++ b/Mathlib/Topology/MetricSpace/Polish.lean @@ -372,7 +372,7 @@ theorem _root_.IsClosed.isClopenable [TopologicalSpace α] [PolishSpace α] {s : simp only [preimage_preimage, f] have inl (x : s) : (Equiv.Set.sumCompl s) (Sum.inl x) = x := Equiv.Set.sumCompl_apply_inl .. have inr (x : ↑sᶜ) : (Equiv.Set.sumCompl s) (Sum.inr x) = x := Equiv.Set.sumCompl_apply_inr .. - simp_rw [inl, inr, Subtype.coe_preimage_self] + simp_rw [t, inl, inr, Subtype.coe_preimage_self] simp only [isOpen_univ, true_and] rw [Subtype.preimage_coe_compl'] simp diff --git a/Mathlib/Topology/UniformSpace/CompactConvergence.lean b/Mathlib/Topology/UniformSpace/CompactConvergence.lean index d5e4df759ebef..e8c974fc11d59 100644 --- a/Mathlib/Topology/UniformSpace/CompactConvergence.lean +++ b/Mathlib/Topology/UniformSpace/CompactConvergence.lean @@ -339,7 +339,7 @@ theorem uniformSpace_eq_inf_precomp_of_cover {δ₁ δ₂ : Type*} [TopologicalS have h_preimage₂ : MapsTo (φ₂ ⁻¹' ·) 𝔖 𝔗₂ := fun K ↦ h_proper₂.isCompact_preimage have h_cover' : ∀ S ∈ 𝔖, S ⊆ range φ₁ ∪ range φ₂ := fun S _ ↦ h_cover ▸ subset_univ _ -- ... and we just pull it back. - simp_rw [compactConvergenceUniformSpace, replaceTopology_eq, + simp_rw +zetaDelta [compactConvergenceUniformSpace, replaceTopology_eq, UniformOnFun.uniformSpace_eq_inf_precomp_of_cover _ _ _ _ _ h_image₁ h_image₂ h_preimage₁ h_preimage₂ h_cover', UniformSpace.comap_inf, ← UniformSpace.comap_comap] @@ -360,7 +360,7 @@ theorem uniformSpace_eq_iInf_precomp_of_cover {δ : ι → Type*} [∀ i, Topolo inter_eq_right.mp ?_⟩ simp_rw [iUnion₂_inter, mem_setOf, iUnion_nonempty_self, ← iUnion_inter, h_cover, univ_inter] -- ... and we just pull it back. - simp_rw [compactConvergenceUniformSpace, replaceTopology_eq, + simp_rw +zetaDelta [compactConvergenceUniformSpace, replaceTopology_eq, UniformOnFun.uniformSpace_eq_iInf_precomp_of_cover _ _ _ h_image h_preimage h_cover', UniformSpace.comap_iInf, ← UniformSpace.comap_comap] rfl From 87843b4fe8504795552f9ebcd77d71f1838a77aa Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Fri, 13 Dec 2024 07:22:38 +0000 Subject: [PATCH 691/829] feat(Algebra/Group/Aut): Automorphism groups of `Multiplicative G` and `Additive G` (#19905) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds isomorphisms `MulAut (Multiplicative G) ≃* AddAut G` and `AddAut (Additive G) ≃* MulAut G`. --- Mathlib/Algebra/Group/Aut.lean | 13 +++++++++++++ 1 file changed, 13 insertions(+) diff --git a/Mathlib/Algebra/Group/Aut.lean b/Mathlib/Algebra/Group/Aut.lean index 012f362054268..a16056a067cb1 100644 --- a/Mathlib/Algebra/Group/Aut.lean +++ b/Mathlib/Algebra/Group/Aut.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Callum Sutton, Yury Kudryashov -/ +import Mathlib.Algebra.Group.Equiv.TypeTags import Mathlib.GroupTheory.Perm.Basic /-! @@ -276,3 +277,15 @@ def congr [AddGroup G] {H : Type*} [AddGroup H] (ϕ : G ≃+ H) : map_mul' := by simp [DFunLike.ext_iff] end AddAut + +variable (G) + +/-- `Multiplicative G` and `G` have isomorphic automorphism groups. -/ +@[simps!] +def MulAutMultiplicative [AddGroup G] : MulAut (Multiplicative G) ≃* AddAut G := + { AddEquiv.toMultiplicative.symm with map_mul' := fun _ _ ↦ rfl } + +/-- `Additive G` and `G` have isomorphic automorphism groups. -/ +@[simps!] +def AddAutAdditive [Group G] : AddAut (Additive G) ≃* MulAut G := + { MulEquiv.toAdditive.symm with map_mul' := fun _ _ ↦ rfl } From f69c743b785034ed9ce0f0fec2234e5e50f0b2e8 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 13 Dec 2024 07:53:59 +0000 Subject: [PATCH 692/829] feat: representatives for the `FixedDetMatrices` under SL action (#16160) Define the set of representatives under the action of SL(2, ZZ) and give some reduction lemmas that are useful for later proving that SL(2, ZZ) is generated by S, T. Co-authored-by: Johan Commelin Co-authored-by: leanprover-community-mathlib4-bot --- .../Matrix/FixedDetMatrices.lean | 200 +++++++++++++++++- .../Matrix/SpecialLinearGroup.lean | 4 + 2 files changed, 200 insertions(+), 4 deletions(-) diff --git a/Mathlib/LinearAlgebra/Matrix/FixedDetMatrices.lean b/Mathlib/LinearAlgebra/Matrix/FixedDetMatrices.lean index 3120568c26c11..0154b1ab6322f 100644 --- a/Mathlib/LinearAlgebra/Matrix/FixedDetMatrices.lean +++ b/Mathlib/LinearAlgebra/Matrix/FixedDetMatrices.lean @@ -5,6 +5,7 @@ Authors: Chris Birkbeck -/ import Mathlib.LinearAlgebra.Matrix.SpecialLinearGroup +import Mathlib.Data.Int.Interval /-! # Matrices with fixed determinant @@ -18,12 +19,13 @@ kbb (https://github.com/kim-em/kbb/tree/master) repository, so credit to those a variable (n : Type*) [DecidableEq n] [Fintype n] (R : Type*) [CommRing R] -/--The set of matrices with fixed determinant `m`. -/ +/--The subtype of matrices with fixed determinant `m`. -/ def FixedDetMatrix (m : R) := { A : Matrix n n R // A.det = m } -namespace FixedDetMatrix +namespace FixedDetMatrices -open ModularGroup Matrix SpecialLinearGroup MatrixGroups +open Matrix hiding mul_smul +open ModularGroup SpecialLinearGroup MatrixGroups /--Extensionality theorem for `FixedDetMatrix` with respect to the underlying matrix, not entriwise. -/ @@ -52,4 +54,194 @@ lemma smul_coe (m : R) (g : SpecialLinearGroup n R) (A : FixedDetMatrix n R m) : (g • A).1 = g * A.1 := by rw [smul_def] -end FixedDetMatrix +section IntegralFixedDetMatrices + +local notation:1024 "Δ" m : 1024 => (FixedDetMatrix (Fin 2) ℤ m) + +variable {m : ℤ} + +/--Set of representatives for the orbits under `S` and `T`. -/ +def reps (m : ℤ) : Set (Δ m) := + {A : Δ m | (A.1 1 0) = 0 ∧ 0 < A.1 0 0 ∧ 0 ≤ A.1 0 1 ∧ |(A.1 0 1)| < |(A.1 1 1)|} + +/--Reduction step for matrices in `Δ m` which moves the matrices towards `reps`.-/ +def reduceStep (A : Δ m) : Δ m := S • (T ^ (-(A.1 0 0 / A.1 1 0))) • A + +private lemma reduce_aux {A : Δ m} (h : (A.1 1 0) ≠ 0) : + |((reduceStep A).1 1 0)| < |(A.1 1 0)| := by + suffices ((reduceStep A).1 1 0) = A.1 0 0 % A.1 1 0 by + rw [this, abs_eq_self.mpr (Int.emod_nonneg (A.1 0 0) h)] + exact Int.emod_lt (A.1 0 0) h + simp_rw [Int.emod_def, sub_eq_add_neg, reduceStep, smul_coe, coe_T_zpow, S] + norm_num [vecMul, vecHead, vecTail, mul_comm] + +/--Reduction lemma for integral FixedDetMatrices. -/ +@[elab_as_elim] +def reduce_rec {C : Δ m → Sort*} + (base : ∀ A : Δ m, (A.1 1 0) = 0 → C A) + (step : ∀ A : Δ m, (A.1 1 0) ≠ 0 → C (reduceStep A) → C A) : + ∀ A, C A := fun A => by + by_cases h : (A.1 1 0) = 0 + · exact base _ h + · exact step A h (reduce_rec base step (reduceStep A)) + termination_by A => Int.natAbs (A.1 1 0) + decreasing_by + zify + exact reduce_aux h + +/--Map from `Δ m → Δ m` which reduces a FixedDetMatrix towards a representative element in reps. -/ +def reduce : Δ m → Δ m := fun A ↦ + if (A.1 1 0) = 0 then + if 0 < A.1 0 0 then (T ^ (-(A.1 0 1 / A.1 1 1))) • A else + (T ^ (-(-A.1 0 1 / -A.1 1 1))) • (S • (S • A)) --the -/- don't cancel with ℤ divs. + else + reduce (reduceStep A) + termination_by b => Int.natAbs (b.1 1 0) + decreasing_by + next a h => + zify + exact reduce_aux h + +lemma reduce_of_pos {A : Δ m} (hc : (A.1 1 0) = 0) (ha : 0 < A.1 0 0) : + reduce A = (T ^ (-(A.1 0 1 / A.1 1 1))) • A := by + rw [reduce] + simp only [zpow_neg, Int.ediv_neg, neg_neg] at * + simp_rw [if_pos hc, if_pos ha] + +lemma reduce_of_not_pos {A : Δ m} (hc : (A.1 1 0) = 0) (ha : ¬ 0 < A.1 0 0) : + reduce A = (T ^ (-(-A.1 0 1 / -A.1 1 1))) • (S • (S • A)) := by + rw [reduce] + simp only [abs_eq_zero, zpow_neg, Int.ediv_neg, neg_neg] at * + simp_rw [if_pos hc, if_neg ha] + +@[simp] +lemma reduce_reduceStep {A : Δ m} (hc : (A.1 1 0) ≠ 0) : + reduce (reduceStep A) = reduce A := by + symm + rw [reduce, if_neg hc] + +private lemma A_c_eq_zero {A : Δ m} (ha : A.1 1 0 = 0) : A.1 0 0 * A.1 1 1 = m := by + simpa only [det_fin_two, ha, mul_zero, sub_zero] using A.2 + +private lemma A_d_ne_zero {A : Δ m} (ha : A.1 1 0 = 0) (hm : m ≠ 0) : A.1 1 1 ≠ 0 := + right_ne_zero_of_mul (A_c_eq_zero (ha) ▸ hm) + +private lemma A_a_ne_zero {A : Δ m} (ha : A.1 1 0 = 0) (hm : m ≠ 0) : A.1 0 0 ≠ 0 := + left_ne_zero_of_mul (A_c_eq_zero ha ▸ hm) + +/--An auxiliary result bounding the size of the entries of the representatives in `reps`. -/ +lemma reps_entries_le_m' {A : Δ m} (h : A ∈ reps m) (i j : Fin 2) : + A.1 i j ∈ Finset.Icc (-|m|) |m| := by + suffices |A.1 i j| ≤ |m| from Finset.mem_Icc.mpr <| abs_le.mp this + obtain ⟨h10, h00, h01, h11⟩ := h + have h1 : 0 < |A.1 1 1| := (abs_nonneg _).trans_lt h11 + have h2 : 0 < |A.1 0 0| := abs_pos.mpr h00.ne' + fin_cases i <;> fin_cases j + · simpa only [← abs_mul, A_c_eq_zero h10] using (le_mul_iff_one_le_right h2).mpr h1 + · simpa only [← abs_mul, A_c_eq_zero h10] using h11.le.trans (le_mul_of_one_le_left h1.le h2) + · simp_all + · simpa only [← abs_mul, A_c_eq_zero h10] using (le_mul_iff_one_le_left h1).mpr h2 + +@[simp] +lemma reps_zero_empty : reps 0 = ∅ := by + rw [reps, Set.eq_empty_iff_forall_not_mem] + rintro A ⟨h₁, h₂, -, h₄⟩ + suffices |A.1 0 1| < 0 by linarith [abs_nonneg (A.1 0 1)] + have := A_c_eq_zero h₁ + simp_all [h₂.ne'] + +noncomputable instance repsFintype (k : ℤ) : Fintype (reps k) := by + let H := Finset.Icc (-|k|) |k| + let H4 := Fin 2 → Fin 2 → H + apply Fintype.ofInjective (β := H4) (f := fun M i j ↦ ⟨M.1.1 i j, reps_entries_le_m' M.2 i j⟩) + intro M N h + ext i j + simpa only [Subtype.mk.injEq] using congrFun₂ h i j + +@[simp] +lemma S_smul_four (A : Δ m) : S • S • S • S • A = A := by + simp only [smul_def, ← mul_assoc, S_mul_S_eq, neg_mul, one_mul, mul_neg, neg_neg, Subtype.coe_eta] + +@[simp] +lemma T_S_rel_smul (A : Δ m) : S • S • S • T • S • T • S • A = T⁻¹ • A := by + simp_rw [← T_S_rel, ← smul_assoc] + +lemma reduce_mem_reps {m : ℤ} (hm : m ≠ 0) (A : Δ m) : reduce A ∈ reps m := by + induction A using reduce_rec with + | step A h1 h2 => simpa only [reduce_reduceStep h1] using h2 + | base A h => + have hd := A_d_ne_zero h hm + by_cases h1 : 0 < A.1 0 0 + · simp only [reduce_of_pos h h1] + have h2 := Int.emod_def (A.1 0 1) (A.1 1 1) + have h4 := Int.ediv_mul_le (A.1 0 1) hd + set n : ℤ := A.1 0 1 / A.1 1 1 + have h3 := Int.emod_lt (A.1 0 1) hd + rw [← abs_eq_self.mpr <| Int.emod_nonneg _ hd] at h3 + simp only [smul_def, Fin.isValue, coe_T_zpow] + suffices A.1 1 0 = 0 ∧ n * A.1 1 0 < A.1 0 0 ∧ + n * A.1 1 1 ≤ A.1 0 1 ∧ |A.1 0 1 + -(n * A.1 1 1)| < |A.1 1 1| by + simpa only [reps, Fin.isValue, cons_mul, Nat.succ_eq_add_one, Nat.reduceAdd, empty_mul, + Equiv.symm_apply_apply, Set.mem_setOf_eq, of_apply, cons_val', vecMul, cons_dotProduct, + vecHead, one_mul, vecTail, Function.comp_apply, Fin.succ_zero_eq_one, neg_mul, + dotProduct_empty, add_zero, zero_mul, zero_add, empty_val', cons_val_fin_one, + cons_val_one, cons_val_zero, lt_add_neg_iff_add_lt, le_add_neg_iff_add_le] + simp_all only [h, mul_comm n, zero_mul, ← sub_eq_add_neg, ← h2, + Fin.isValue, h1, h3, and_true, true_and] + · simp only [reps, Fin.isValue, reduce_of_not_pos h h1, Int.ediv_neg, neg_neg, smul_def, ← + mul_assoc, S_mul_S_eq, neg_mul, one_mul, coe_T_zpow, mul_neg, cons_mul, Nat.succ_eq_add_one, + Nat.reduceAdd, empty_mul, Equiv.symm_apply_apply, neg_of, neg_cons, neg_empty, + Set.mem_setOf_eq, of_apply, cons_val', Pi.neg_apply, vecMul, cons_dotProduct, vecHead, + vecTail, Function.comp_apply, Fin.succ_zero_eq_one, h, mul_zero, dotProduct_empty, add_zero, + zero_mul, neg_zero, empty_val', cons_val_fin_one, cons_val_one, cons_val_zero, lt_neg, + neg_add_rev, zero_add, le_add_neg_iff_add_le, ← le_neg, abs_neg, true_and] + refine ⟨?_, Int.ediv_mul_le _ hd, ?_⟩ + · simp only [Int.lt_iff_le_and_ne] + exact ⟨not_lt.mp h1, A_a_ne_zero h hm⟩ + · rw [mul_comm, add_comm, ← Int.sub_eq_add_neg, ← Int.emod_def, + abs_eq_self.mpr <| Int.emod_nonneg _ hd] + exact Int.emod_lt _ hd + +variable {C : Δ m → Prop} + +private lemma prop_red_S (hS : ∀ B, C B → C (S • B)) (B) : C (S • B) ↔ C B := by + refine ⟨?_, hS _⟩ + intro ih + rw [← (S_smul_four B)] + solve_by_elim + +private lemma prop_red_T (hS : ∀ B, C B → C (S • B)) (hT : ∀ B, C B → C (T • B)) (B) : + C (T • B) ↔ C B := by + refine ⟨?_, hT _⟩ + intro ih + rw [show B = T⁻¹ • T • B by simp, ← T_S_rel_smul] + solve_by_elim (config := {maxDepth := 10}) + +private lemma prop_red_T_pow (hS : ∀ B, C B → C (S • B)) (hT : ∀ B, C B → C (T • B)) : + ∀ B (n : ℤ), C (T^n • B) ↔ C B := by + intro B n + induction' n using Int.induction_on with n hn m hm + · simp only [zpow_zero, one_smul, imp_self] + · simpa only [add_comm (n:ℤ), zpow_add _ 1, ← smul_eq_mul, zpow_one, smul_assoc, prop_red_T hS hT] + · rwa [sub_eq_neg_add, zpow_add, zpow_neg_one, ← prop_red_T hS hT, mul_smul, smul_inv_smul] + +@[elab_as_elim] +theorem induction_on {C : Δ m → Prop} {A : Δ m} (hm : m ≠ 0) + (h0 : ∀ A : Δ m, A.1 1 0 = 0 → 0 < A.1 0 0 → 0 ≤ A.1 0 1 → |(A.1 0 1)| < |(A.1 1 1)| → C A) + (hS : ∀ B, C B → C (S • B)) (hT : ∀ B, C B → C (T • B)) : C A := by + have h_reduce : C (reduce A) := by + rcases reduce_mem_reps hm A with ⟨H1, H2, H3, H4⟩ + exact h0 _ H1 H2 H3 H4 + suffices ∀ A : Δ m, C (reduce A) → C A from this _ h_reduce + apply reduce_rec + · intro A h + by_cases h1 : 0 < A.1 0 0 + · simp only [reduce_of_pos h h1, prop_red_T_pow hS hT, imp_self] + · simp only [reduce_of_not_pos h h1, prop_red_T_pow hS hT, prop_red_S hS, imp_self] + intro A hc ih hA + rw [← reduce_reduceStep hc] at hA + simpa only [reduceStep, prop_red_S hS, prop_red_T_pow hS hT] using ih hA + +end IntegralFixedDetMatrices + +end FixedDetMatrices diff --git a/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean b/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean index 59bff4a8773bd..a3dfd72c6e079 100644 --- a/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean +++ b/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean @@ -491,4 +491,8 @@ lemma S_mul_S_eq : (S : Matrix (Fin 2) (Fin 2) ℤ) * S = -1 := by empty_vecMul, add_zero, zero_add, empty_mul, Equiv.symm_apply_apply] exact Eq.symm (eta_fin_two (-1)) +lemma T_S_rel : S • S • S • T • S • T • S = T⁻¹ := by + ext i j + fin_cases i <;> fin_cases j <;> rfl + end ModularGroup From 383b0748ba499f742d97bf6ff772502b1e458986 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Fri, 13 Dec 2024 07:54:01 +0000 Subject: [PATCH 693/829] feat(CategoryTheory/Localization): liftings of bifunctors (#19894) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit When a functor `F : C ⥤ E` inverts a class of morphisms `W : MorphismProperty C` and `L : C ⥤ D` is a localization functor for `W`, then `F` lifts as a functor `D ⥤ E`. In this PR, this is generalized to bifunctors (or "functors in two variables") `F : C₁ ⥤ C₂ ⥤ E`, which may be lifted as a functor `D₁ ⥤ D₂ ⥤ E` under similar assumptions. (We shall also need a version of this in three variables.) --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Functor/Category.lean | 12 ++ Mathlib/CategoryTheory/Functor/Currying.lean | 20 ++ .../Localization/Bifunctor.lean | 178 ++++++++++++++++++ .../Localization/Predicate.lean | 6 +- Mathlib/CategoryTheory/NatIso.lean | 10 + Mathlib/CategoryTheory/Whiskering.lean | 16 ++ 7 files changed, 240 insertions(+), 3 deletions(-) create mode 100644 Mathlib/CategoryTheory/Localization/Bifunctor.lean diff --git a/Mathlib.lean b/Mathlib.lean index 5845af9e4e140..c7c126515b79a 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1923,6 +1923,7 @@ import Mathlib.CategoryTheory.Linear.FunctorCategory import Mathlib.CategoryTheory.Linear.LinearFunctor import Mathlib.CategoryTheory.Linear.Yoneda import Mathlib.CategoryTheory.Localization.Adjunction +import Mathlib.CategoryTheory.Localization.Bifunctor import Mathlib.CategoryTheory.Localization.Bousfield import Mathlib.CategoryTheory.Localization.CalculusOfFractions import Mathlib.CategoryTheory.Localization.CalculusOfFractions.ComposableArrows diff --git a/Mathlib/CategoryTheory/Functor/Category.lean b/Mathlib/CategoryTheory/Functor/Category.lean index 7288ae9ebc9fe..64d1a3f601533 100644 --- a/Mathlib/CategoryTheory/Functor/Category.lean +++ b/Mathlib/CategoryTheory/Functor/Category.lean @@ -142,6 +142,18 @@ protected def flip (F : C ⥤ D ⥤ E) : D ⥤ C ⥤ E where end Functor +variable (C D E) in +/-- The functor `(C ⥤ D ⥤ E) ⥤ D ⥤ C ⥤ E` which flips the variables. -/ +@[simps] +def flipFunctor : (C ⥤ D ⥤ E) ⥤ D ⥤ C ⥤ E where + obj F := F.flip + map {F₁ F₂} φ := + { app := fun Y => + { app := fun X => (φ.app X).app Y + naturality := fun X₁ X₂ f => by + dsimp + simp only [← NatTrans.comp_app, naturality] } } + namespace Iso @[reassoc (attr := simp)] diff --git a/Mathlib/CategoryTheory/Functor/Currying.lean b/Mathlib/CategoryTheory/Functor/Currying.lean index 5c8021e9a84c0..67683abeb8e5a 100644 --- a/Mathlib/CategoryTheory/Functor/Currying.lean +++ b/Mathlib/CategoryTheory/Functor/Currying.lean @@ -88,6 +88,26 @@ def currying : C ⥤ D ⥤ E ≌ C × D ⥤ E where dsimp at f₁ f₂ ⊢ simp only [← F.map_comp, prod_comp, Category.comp_id, Category.id_comp])) +/-- The functor `uncurry : (C ⥤ D ⥤ E) ⥤ C × D ⥤ E` is fully faithful. -/ +def fullyFaithfulUncurry : (uncurry : (C ⥤ D ⥤ E) ⥤ C × D ⥤ E).FullyFaithful := + currying.fullyFaithfulFunctor + +instance : (uncurry : (C ⥤ D ⥤ E) ⥤ C × D ⥤ E).Full := + fullyFaithfulUncurry.full + +instance : (uncurry : (C ⥤ D ⥤ E) ⥤ C × D ⥤ E).Faithful := + fullyFaithfulUncurry.faithful + +/-- Given functors `F₁ : C ⥤ D`, `F₂ : C' ⥤ D'` and `G : D × D' ⥤ E`, this is the isomorphism +between `curry.obj ((F₁.prod F₂).comp G)` and +`F₁ ⋙ curry.obj G ⋙ (whiskeringLeft C' D' E).obj F₂` in the category `C ⥤ C' ⥤ E`. -/ +@[simps!] +def curryObjProdComp {C' D' : Type*} [Category C'] [Category D'] + (F₁ : C ⥤ D) (F₂ : C' ⥤ D') (G : D × D' ⥤ E) : + curry.obj ((F₁.prod F₂).comp G) ≅ + F₁ ⋙ curry.obj G ⋙ (whiskeringLeft C' D' E).obj F₂ := + NatIso.ofComponents (fun X₁ ↦ NatIso.ofComponents (fun X₂ ↦ Iso.refl _)) + /-- `F.flip` is isomorphic to uncurrying `F`, swapping the variables, and currying. -/ @[simps!] def flipIsoCurrySwapUncurry (F : C ⥤ D ⥤ E) : F.flip ≅ curry.obj (Prod.swap _ _ ⋙ uncurry.obj F) := diff --git a/Mathlib/CategoryTheory/Localization/Bifunctor.lean b/Mathlib/CategoryTheory/Localization/Bifunctor.lean new file mode 100644 index 0000000000000..e4556b6e608d8 --- /dev/null +++ b/Mathlib/CategoryTheory/Localization/Bifunctor.lean @@ -0,0 +1,178 @@ +/- +Copyright (c) 2023 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.CategoryTheory.Localization.Prod +import Mathlib.CategoryTheory.Functor.Currying + +/-! +# Lifting of bifunctors + +In this file, in the context of the localization of categories, we extend the notion +of lifting of functors to the case of bifunctors. As the localization of categories +behaves well with respect to finite products of categories (when the classes of +morphisms contain identities), all the definitions for bifunctors `C₁ ⥤ C₂ ⥤ E` +are obtained by reducing to the case of functors `(C₁ × C₂) ⥤ E` by using +currying and uncurrying. + +Given morphism properties `W₁ : MorphismProperty C₁` and `W₂ : MorphismProperty C₂`, +and a functor `F : C₁ ⥤ C₂ ⥤ E`, we define `MorphismProperty.IsInvertedBy₂ W₁ W₂ F` +as the condition that the functor `uncurry.obj F : C₁ × C₂ ⥤ E` inverts `W₁.prod W₂`. + +If `L₁ : C₁ ⥤ D₁` and `L₂ : C₂ ⥤ D₂` are localization functors for `W₁` and `W₂` +respectively, and `F : C₁ ⥤ C₂ ⥤ E` satisfies `MorphismProperty.IsInvertedBy₂ W₁ W₂ F`, +we introduce `Localization.lift₂ F hF L₁ L₂ : D₁ ⥤ D₂ ⥤ E` which is a bifunctor +which lifts `F`. + +-/ + +namespace CategoryTheory + +open Category + +variable {C₁ C₂ D₁ D₂ E E' : Type*} [Category C₁] [Category C₂] + [Category D₁] [Category D₂] [Category E] [Category E'] + +namespace MorphismProperty + +/-- Classes of morphisms `W₁ : MorphismProperty C₁` and `W₂ : MorphismProperty C₂` are said +to be inverted by `F : C₁ ⥤ C₂ ⥤ E` if `W₁.prod W₂` is inverted by +the functor `uncurry.obj F : C₁ × C₂ ⥤ E`. -/ +def IsInvertedBy₂ (W₁ : MorphismProperty C₁) (W₂ : MorphismProperty C₂) + (F : C₁ ⥤ C₂ ⥤ E) : Prop := + (W₁.prod W₂).IsInvertedBy (uncurry.obj F) + +end MorphismProperty + +namespace Localization + +section + +variable (L₁ : C₁ ⥤ D₁) (L₂ : C₂ ⥤ D₂) + +/-- Given functors `L₁ : C₁ ⥤ D₁`, `L₂ : C₂ ⥤ D₂`, morphisms properties `W₁` on `C₁` +and `W₂` on `C₂`, and functors `F : C₁ ⥤ C₂ ⥤ E` and `F' : D₁ ⥤ D₂ ⥤ E`, we say +`Lifting₂ L₁ L₂ W₁ W₂ F F'` holds if `F` is induced by `F'`, up to an isomorphism. -/ +class Lifting₂ (W₁ : MorphismProperty C₁) (W₂ : MorphismProperty C₂) + (F : C₁ ⥤ C₂ ⥤ E) (F' : D₁ ⥤ D₂ ⥤ E) where + /-- the isomorphism `(((whiskeringLeft₂ E).obj L₁).obj L₂).obj F' ≅ F` expressing + that `F` is induced by `F'` up to an isomorphism -/ + iso' : (((whiskeringLeft₂ E).obj L₁).obj L₂).obj F' ≅ F + +variable (W₁ : MorphismProperty C₁) (W₂ : MorphismProperty C₂) + (F : C₁ ⥤ C₂ ⥤ E) (F' : D₁ ⥤ D₂ ⥤ E) [Lifting₂ L₁ L₂ W₁ W₂ F F'] + +/-- The isomorphism `(((whiskeringLeft₂ E).obj L₁).obj L₂).obj F' ≅ F` when +`Lifting₂ L₁ L₂ W₁ W₂ F F'` holds. -/ +noncomputable def Lifting₂.iso : (((whiskeringLeft₂ E).obj L₁).obj L₂).obj F' ≅ F := + Lifting₂.iso' W₁ W₂ + +/-- If `Lifting₂ L₁ L₂ W₁ W₂ F F'` holds, then `Lifting L₂ W₂ (F.obj X₁) (F'.obj (L₁.obj X₁))` +holds for any `X₁ : C₁`. -/ +noncomputable def Lifting₂.fst (X₁ : C₁) : + Lifting L₂ W₂ (F.obj X₁) (F'.obj (L₁.obj X₁)) where + iso' := ((evaluation _ _).obj X₁).mapIso (Lifting₂.iso L₁ L₂ W₁ W₂ F F') + +noncomputable instance Lifting₂.flip : Lifting₂ L₂ L₁ W₂ W₁ F.flip F'.flip where + iso' := (flipFunctor _ _ _).mapIso (Lifting₂.iso L₁ L₂ W₁ W₂ F F') + +/-- If `Lifting₂ L₁ L₂ W₁ W₂ F F'` holds, then +`Lifting L₁ W₁ (F.flip.obj X₂) (F'.flip.obj (L₂.obj X₂))` holds for any `X₂ : C₂`. -/ +noncomputable def Lifting₂.snd (X₂ : C₂) : + Lifting L₁ W₁ (F.flip.obj X₂) (F'.flip.obj (L₂.obj X₂)) := + Lifting₂.fst L₂ L₁ W₂ W₁ F.flip F'.flip X₂ + +noncomputable instance Lifting₂.uncurry [Lifting₂ L₁ L₂ W₁ W₂ F F'] : + Lifting (L₁.prod L₂) (W₁.prod W₂) (uncurry.obj F) (uncurry.obj F') where + iso' := CategoryTheory.uncurry.mapIso (Lifting₂.iso L₁ L₂ W₁ W₂ F F') + +end + +section + +variable (F : C₁ ⥤ C₂ ⥤ E) {W₁ : MorphismProperty C₁} {W₂ : MorphismProperty C₂} + (hF : MorphismProperty.IsInvertedBy₂ W₁ W₂ F) + (L₁ : C₁ ⥤ D₁) (L₂ : C₂ ⥤ D₂) + [L₁.IsLocalization W₁] [L₂.IsLocalization W₂] + [W₁.ContainsIdentities] [W₂.ContainsIdentities] + +/-- Given localization functor `L₁ : C₁ ⥤ D₁` and `L₂ : C₂ ⥤ D₂` with respect +to `W₁ : MorphismProperty C₁` and `W₂ : MorphismProperty C₂` respectively, +and a bifunctor `F : C₁ ⥤ C₂ ⥤ E` which inverts `W₁` and `W₂`, this is +the induced localized bifunctor `D₁ ⥤ D₂ ⥤ E`. -/ +noncomputable def lift₂ : D₁ ⥤ D₂ ⥤ E := + curry.obj (lift (uncurry.obj F) hF (L₁.prod L₂)) + +noncomputable instance : Lifting₂ L₁ L₂ W₁ W₂ F (lift₂ F hF L₁ L₂) where + iso' := (curryObjProdComp _ _ _).symm ≪≫ + curry.mapIso (fac (uncurry.obj F) hF (L₁.prod L₂)) ≪≫ + currying.unitIso.symm.app F + +noncomputable instance Lifting₂.liftingLift₂ (X₁ : C₁) : + Lifting L₂ W₂ (F.obj X₁) ((lift₂ F hF L₁ L₂).obj (L₁.obj X₁)) := + Lifting₂.fst _ _ W₁ _ _ _ _ + +noncomputable instance Lifting₂.liftingLift₂Flip (X₂ : C₂) : + Lifting L₁ W₁ (F.flip.obj X₂) ((lift₂ F hF L₁ L₂).flip.obj (L₂.obj X₂)) := + Lifting₂.snd _ _ _ W₂ _ _ _ + +lemma lift₂_iso_hom_app_app₁ (X₁ : C₁) (X₂ : C₂) : + ((Lifting₂.iso L₁ L₂ W₁ W₂ F (lift₂ F hF L₁ L₂)).hom.app X₁).app X₂ = + (Lifting.iso L₂ W₂ (F.obj X₁) ((lift₂ F hF L₁ L₂).obj (L₁.obj X₁))).hom.app X₂ := + rfl + +lemma lift₂_iso_hom_app_app₂ (X₁ : C₁) (X₂ : C₂) : + ((Lifting₂.iso L₁ L₂ W₁ W₂ F (lift₂ F hF L₁ L₂)).hom.app X₁).app X₂ = + (Lifting.iso L₁ W₁ (F.flip.obj X₂) ((lift₂ F hF L₁ L₂).flip.obj (L₂.obj X₂))).hom.app X₁ := + rfl + +end + +section + +variable (L₁ : C₁ ⥤ D₁) (L₂ : C₂ ⥤ D₂) + (W₁ : MorphismProperty C₁) (W₂ : MorphismProperty C₂) + [L₁.IsLocalization W₁] [L₂.IsLocalization W₂] + [W₁.ContainsIdentities] [W₂.ContainsIdentities] + (F₁ F₂ : C₁ ⥤ C₂ ⥤ E) (F₁' F₂' : D₁ ⥤ D₂ ⥤ E) + [Lifting₂ L₁ L₂ W₁ W₂ F₁ F₁'] [Lifting₂ L₁ L₂ W₁ W₂ F₂ F₂'] + +/-- The natural transformation `F₁' ⟶ F₂'` of bifunctors induced by a +natural transformation `τ : F₁ ⟶ F₂` when `Lifting₂ L₁ L₂ W₁ W₂ F₁ F₁'` +and `Lifting₂ L₁ L₂ W₁ W₂ F₂ F₂'` hold. -/ +noncomputable def lift₂NatTrans (τ : F₁ ⟶ F₂) : F₁' ⟶ F₂' := + fullyFaithfulUncurry.preimage + (liftNatTrans (L₁.prod L₂) (W₁.prod W₂) (uncurry.obj F₁) + (uncurry.obj F₂) (uncurry.obj F₁') (uncurry.obj F₂') (uncurry.map τ)) + +@[simp] +theorem lift₂NatTrans_app_app (τ : F₁ ⟶ F₂) (X₁ : C₁) (X₂ : C₂) : + ((lift₂NatTrans L₁ L₂ W₁ W₂ F₁ F₂ F₁' F₂' τ).app (L₁.obj X₁)).app (L₂.obj X₂) = + ((Lifting₂.iso L₁ L₂ W₁ W₂ F₁ F₁').hom.app X₁).app X₂ ≫ (τ.app X₁).app X₂ ≫ + ((Lifting₂.iso L₁ L₂ W₁ W₂ F₂ F₂').inv.app X₁).app X₂ := by + dsimp [lift₂NatTrans, fullyFaithfulUncurry, Equivalence.fullyFaithfulFunctor] + simp only [currying_unitIso_hom_app_app_app, currying_unitIso_inv_app_app_app, comp_id, id_comp] + exact liftNatTrans_app _ _ _ _ (uncurry.obj F₁') (uncurry.obj F₂') (uncurry.map τ) ⟨X₁, X₂⟩ + +variable {F₁' F₂'} in +include W₁ W₂ in +theorem natTrans₂_ext {τ τ' : F₁' ⟶ F₂'} + (h : ∀ (X₁ : C₁) (X₂ : C₂), (τ.app (L₁.obj X₁)).app (L₂.obj X₂) = + (τ'.app (L₁.obj X₁)).app (L₂.obj X₂)) : τ = τ' := + uncurry.map_injective (natTrans_ext (L₁.prod L₂) (W₁.prod W₂) (fun _ ↦ h _ _)) + +/-- The natural isomorphism `F₁' ≅ F₂'` of bifunctors induced by a +natural isomorphism `e : F₁ ≅ F₂` when `Lifting₂ L₁ L₂ W₁ W₂ F₁ F₁'` +and `Lifting₂ L₁ L₂ W₁ W₂ F₂ F₂'` hold. -/ +noncomputable def lift₂NatIso (e : F₁ ≅ F₂) : F₁' ≅ F₂' where + hom := lift₂NatTrans L₁ L₂ W₁ W₂ F₁ F₂ F₁' F₂' e.hom + inv := lift₂NatTrans L₁ L₂ W₁ W₂ F₂ F₁ F₂' F₁' e.inv + hom_inv_id := natTrans₂_ext L₁ L₂ W₁ W₂ (by aesop_cat) + inv_hom_id := natTrans₂_ext L₁ L₂ W₁ W₂ (by aesop_cat) + +end + +end Localization + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Localization/Predicate.lean b/Mathlib/CategoryTheory/Localization/Predicate.lean index b56d44cf24b6c..bc2855fa278ac 100644 --- a/Mathlib/CategoryTheory/Localization/Predicate.lean +++ b/Mathlib/CategoryTheory/Localization/Predicate.lean @@ -266,7 +266,7 @@ lemma faithful_whiskeringLeft (L : C ⥤ D) (W) [L.IsLocalization W] (E : Type*) variable {E} -theorem natTrans_ext (L : C ⥤ D) (W) [L.IsLocalization W] {F₁ F₂ : D ⥤ E} (τ τ' : F₁ ⟶ F₂) +theorem natTrans_ext (L : C ⥤ D) (W) [L.IsLocalization W] {F₁ F₂ : D ⥤ E} {τ τ' : F₁ ⟶ F₂} (h : ∀ X : C, τ.app (L.obj X) = τ'.app (L.obj X)) : τ = τ' := by haveI := essSurj L W ext Y @@ -331,13 +331,13 @@ theorem comp_liftNatTrans (F₁ F₂ F₃ : C ⥤ E) (F₁' F₂' F₃' : D ⥤ [h₂ : Lifting L W F₂ F₂'] [h₃ : Lifting L W F₃ F₃'] (τ : F₁ ⟶ F₂) (τ' : F₂ ⟶ F₃) : liftNatTrans L W F₁ F₂ F₁' F₂' τ ≫ liftNatTrans L W F₂ F₃ F₂' F₃' τ' = liftNatTrans L W F₁ F₃ F₁' F₃' (τ ≫ τ') := - natTrans_ext L W _ _ fun X => by + natTrans_ext L W fun X => by simp only [NatTrans.comp_app, liftNatTrans_app, assoc, Iso.inv_hom_id_app_assoc] @[simp] theorem liftNatTrans_id (F : C ⥤ E) (F' : D ⥤ E) [h : Lifting L W F F'] : liftNatTrans L W F F F' F' (𝟙 F) = 𝟙 F' := - natTrans_ext L W _ _ fun X => by + natTrans_ext L W fun X => by simp only [liftNatTrans_app, NatTrans.id_app, id_comp, Iso.hom_inv_id_app] rfl diff --git a/Mathlib/CategoryTheory/NatIso.lean b/Mathlib/CategoryTheory/NatIso.lean index 2899c82c95db6..fcf0a2f963969 100644 --- a/Mathlib/CategoryTheory/NatIso.lean +++ b/Mathlib/CategoryTheory/NatIso.lean @@ -65,6 +65,16 @@ theorem inv_hom_id_app {F G : C ⥤ D} (α : F ≅ G) (X : C) : α.inv.app X ≫ α.hom.app X = 𝟙 (G.obj X) := congr_fun (congr_arg NatTrans.app α.inv_hom_id) X +@[reassoc (attr := simp)] +lemma hom_inv_id_app_app {F G : C ⥤ D ⥤ E} (e : F ≅ G) (X₁ : C) (X₂ : D) : + (e.hom.app X₁).app X₂ ≫ (e.inv.app X₁).app X₂ = 𝟙 _ := by + rw [← NatTrans.comp_app, Iso.hom_inv_id_app, NatTrans.id_app] + +@[reassoc (attr := simp)] +lemma inv_hom_id_app_app {F G : C ⥤ D ⥤ E} (e : F ≅ G) (X₁ : C) (X₂ : D) : + (e.inv.app X₁).app X₂ ≫ (e.hom.app X₁).app X₂ = 𝟙 _ := by + rw [← NatTrans.comp_app, Iso.inv_hom_id_app, NatTrans.id_app] + end Iso namespace NatIso diff --git a/Mathlib/CategoryTheory/Whiskering.lean b/Mathlib/CategoryTheory/Whiskering.lean index 50cc4f617eec0..53feb19a7b2f9 100644 --- a/Mathlib/CategoryTheory/Whiskering.lean +++ b/Mathlib/CategoryTheory/Whiskering.lean @@ -302,4 +302,20 @@ theorem pentagon : end Functor +variable {C₁ C₂ D₁ D₂ : Type*} [Category C₁] [Category C₂] + [Category D₁] [Category D₂] (E : Type*) [Category E] + +/-- The obvious functor `(C₁ ⥤ D₁) ⥤ (C₂ ⥤ D₂) ⥤ (D₁ ⥤ D₂ ⥤ E) ⥤ (C₁ ⥤ C₂ ⥤ E)`. -/ +@[simps!] +def whiskeringLeft₂ : + (C₁ ⥤ D₁) ⥤ (C₂ ⥤ D₂) ⥤ (D₁ ⥤ D₂ ⥤ E) ⥤ (C₁ ⥤ C₂ ⥤ E) where + obj F₁ := + { obj := fun F₂ ↦ + (whiskeringRight D₁ (D₂ ⥤ E) (C₂ ⥤ E)).obj ((whiskeringLeft C₂ D₂ E).obj F₂) ⋙ + (whiskeringLeft C₁ D₁ (C₂ ⥤ E)).obj F₁ + map := fun φ ↦ whiskerRight + ((whiskeringRight D₁ (D₂ ⥤ E) (C₂ ⥤ E)).map ((whiskeringLeft C₂ D₂ E).map φ)) _ } + map ψ := + { app := fun F₂ ↦ whiskerLeft _ ((whiskeringLeft C₁ D₁ (C₂ ⥤ E)).map ψ) } + end CategoryTheory From 2d8f64e7a9340bd3d4bce48d9498f903d01884de Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 13 Dec 2024 08:43:41 +0000 Subject: [PATCH 694/829] feat: generalize `LinearMap.det_smul` (#19885) --- Mathlib/LinearAlgebra/Determinant.lean | 23 ++++++++++--------- .../LinearAlgebra/FreeModule/Determinant.lean | 2 +- Mathlib/MeasureTheory/Function/Jacobian.lean | 3 ++- 3 files changed, 15 insertions(+), 13 deletions(-) diff --git a/Mathlib/LinearAlgebra/Determinant.lean b/Mathlib/LinearAlgebra/Determinant.lean index 66160703d9068..f1f5ef6f34822 100644 --- a/Mathlib/LinearAlgebra/Determinant.lean +++ b/Mathlib/LinearAlgebra/Determinant.lean @@ -3,6 +3,7 @@ Copyright (c) 2019 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Patrick Massot, Casper Putz, Anne Baanen -/ +import Mathlib.LinearAlgebra.FreeModule.StrongRankCondition import Mathlib.LinearAlgebra.GeneralLinearGroup import Mathlib.LinearAlgebra.Matrix.Reindex import Mathlib.Tactic.FieldSimp @@ -229,17 +230,17 @@ theorem det_id : LinearMap.det (LinearMap.id : M →ₗ[A] M) = 1 := /-- Multiplying a map by a scalar `c` multiplies its determinant by `c ^ dim M`. -/ @[simp] -theorem det_smul {𝕜 : Type*} [Field 𝕜] {M : Type*} [AddCommGroup M] [Module 𝕜 M] (c : 𝕜) - (f : M →ₗ[𝕜] M) : - 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 +theorem det_smul [Module.Free A M] (c : A) (f : M →ₗ[A] M) : + LinearMap.det (c • f) = c ^ Module.finrank A M * LinearMap.det f := by + nontriviality A + by_cases H : ∃ s : Finset M, Nonempty (Basis s A M) + · have : Module.Finite A M := by rcases H with ⟨s, ⟨hs⟩⟩ - exact FiniteDimensional.of_fintype_basis hs - simp only [← det_toMatrix (Module.finBasis 𝕜 M), LinearEquiv.map_smul, + exact Module.Finite.of_basis hs + simp only [← det_toMatrix (Module.finBasis A M), LinearEquiv.map_smul, Fintype.card_fin, Matrix.det_smul] · classical - have : Module.finrank 𝕜 M = 0 := finrank_eq_zero_of_not_exists_basis H + have : Module.finrank A 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) : @@ -252,9 +253,9 @@ theorem det_zero' {ι : Type*} [Finite ι] [Nonempty ι] (b : Basis ι A M) : and `0` otherwise. We give a formula that also works in infinite dimension, where we define the determinant to be `1`. -/ @[simp] -theorem det_zero {𝕜 : Type*} [Field 𝕜] {M : Type*} [AddCommGroup M] [Module 𝕜 M] : - 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_zero [Module.Free A M] : + LinearMap.det (0 : M →ₗ[A] M) = (0 : A) ^ Module.finrank A M := by + simp only [← zero_smul A (1 : M →ₗ[A] M), det_smul, mul_one, MonoidHom.map_one] theorem det_eq_one_of_subsingleton [Subsingleton M] (f : M →ₗ[R] M) : LinearMap.det (f : M →ₗ[R] M) = 1 := by diff --git a/Mathlib/LinearAlgebra/FreeModule/Determinant.lean b/Mathlib/LinearAlgebra/FreeModule/Determinant.lean index 5ff0195e213ee..f7a7432ff9142 100644 --- a/Mathlib/LinearAlgebra/FreeModule/Determinant.lean +++ b/Mathlib/LinearAlgebra/FreeModule/Determinant.lean @@ -19,7 +19,7 @@ free (finite) modules over any commutative ring. -/ -@[simp] +@[simp high] theorem LinearMap.det_zero'' {R M : Type*} [CommRing R] [AddCommGroup M] [Module R M] [Module.Free R M] [Module.Finite R M] [Nontrivial M] : LinearMap.det (0 : M →ₗ[R] M) = 0 := by letI : Nonempty (Module.Free.ChooseBasisIndex R M) := (Module.Free.chooseBasis R M).index_nonempty diff --git a/Mathlib/MeasureTheory/Function/Jacobian.lean b/Mathlib/MeasureTheory/Function/Jacobian.lean index 77a1878659c77..f4914c6260a11 100644 --- a/Mathlib/MeasureTheory/Function/Jacobian.lean +++ b/Mathlib/MeasureTheory/Function/Jacobian.lean @@ -1177,8 +1177,9 @@ theorem integral_image_eq_integral_abs_det_fderiv_smul (hs : MeasurableSet s) rw [NNReal.smul_def, Real.coe_toNNReal _ (abs_nonneg (f' x).det)] -- Porting note: move this to `Topology.Algebra.Module.Basic` when port is over -theorem det_one_smulRight {𝕜 : Type*} [NormedField 𝕜] (v : 𝕜) : +theorem det_one_smulRight {𝕜 : Type*} [CommRing 𝕜] [TopologicalSpace 𝕜] [ContinuousMul 𝕜] (v : 𝕜) : ((1 : 𝕜 →L[𝕜] 𝕜).smulRight v).det = v := by + nontriviality 𝕜 have : (1 : 𝕜 →L[𝕜] 𝕜).smulRight v = v • (1 : 𝕜 →L[𝕜] 𝕜) := by ext1 simp only [ContinuousLinearMap.smulRight_apply, ContinuousLinearMap.one_apply, From e42ca30b426589436293e6f614df7349f9c3b043 Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Fri, 13 Dec 2024 08:59:38 +0000 Subject: [PATCH 695/829] feat(CategoryTheory/Abelian): AB axioms carry over to functor categories (#19914) --- Mathlib.lean | 4 +- Mathlib/Algebra/Category/Grp/AB.lean | 2 +- .../Basic.lean} | 44 ++++++++++++++++++- .../GrothendieckAxioms/FunctorCategory.lean | 29 ++++++++++++ .../Abelian/GrothendieckAxioms/Sheaf.lean | 32 ++++++++++++++ Mathlib/Condensed/Light/AB.lean | 3 +- 6 files changed, 110 insertions(+), 4 deletions(-) rename Mathlib/CategoryTheory/Abelian/{GrothendieckAxioms.lean => GrothendieckAxioms/Basic.lean} (89%) create mode 100644 Mathlib/CategoryTheory/Abelian/GrothendieckAxioms/FunctorCategory.lean create mode 100644 Mathlib/CategoryTheory/Abelian/GrothendieckAxioms/Sheaf.lean diff --git a/Mathlib.lean b/Mathlib.lean index c7c126515b79a..c63dcf2117b60 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1569,7 +1569,9 @@ import Mathlib.CategoryTheory.Abelian.Exact import Mathlib.CategoryTheory.Abelian.Ext import Mathlib.CategoryTheory.Abelian.FunctorCategory import Mathlib.CategoryTheory.Abelian.Generator -import Mathlib.CategoryTheory.Abelian.GrothendieckAxioms +import Mathlib.CategoryTheory.Abelian.GrothendieckAxioms.Basic +import Mathlib.CategoryTheory.Abelian.GrothendieckAxioms.FunctorCategory +import Mathlib.CategoryTheory.Abelian.GrothendieckAxioms.Sheaf import Mathlib.CategoryTheory.Abelian.Images import Mathlib.CategoryTheory.Abelian.Injective import Mathlib.CategoryTheory.Abelian.InjectiveResolution diff --git a/Mathlib/Algebra/Category/Grp/AB.lean b/Mathlib/Algebra/Category/Grp/AB.lean index edbbcb51bdb57..7731a8f49c9fc 100644 --- a/Mathlib/Algebra/Category/Grp/AB.lean +++ b/Mathlib/Algebra/Category/Grp/AB.lean @@ -6,7 +6,7 @@ Authors: David Kurniadi Angdinata, Moritz Firsching, Nikolas Kuhn, Amelia Living import Mathlib.Algebra.Category.Grp.Biproducts import Mathlib.Algebra.Category.Grp.FilteredColimits import Mathlib.Algebra.Homology.ShortComplex.Ab -import Mathlib.CategoryTheory.Abelian.GrothendieckAxioms +import Mathlib.CategoryTheory.Abelian.GrothendieckAxioms.Basic import Mathlib.CategoryTheory.Limits.FunctorCategory.EpiMono /-! # AB axioms for the category of abelian groups diff --git a/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms.lean b/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms/Basic.lean similarity index 89% rename from Mathlib/CategoryTheory/Abelian/GrothendieckAxioms.lean rename to Mathlib/CategoryTheory/Abelian/GrothendieckAxioms/Basic.lean index 9e81a28ffd8e2..251628d72e070 100644 --- a/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms.lean +++ b/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms/Basic.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2023 Adam Topaz. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Isaac Hernando, Coleton Kotch, Adam Topaz +Authors: Dagur Asgeirsson, Isaac Hernando, Coleton Kotch, Adam Topaz -/ import Mathlib.Algebra.Homology.ShortComplex.ExactFunctor import Mathlib.CategoryTheory.Abelian.FunctorCategory @@ -79,6 +79,48 @@ class HasExactLimitsOfShape (J : Type u') [Category.{v'} J] (C : Type u) [Catego attribute [instance] HasExactColimitsOfShape.preservesFiniteLimits HasExactLimitsOfShape.preservesFiniteColimits +variable {C} in +/-- +Pull back a `HasExactColimitsOfShape J` along a functor which preserves and reflects finite limits +and preserves colimits of shape `J` +-/ +lemma HasExactColimitsOfShape.domain_of_functor {D : Type*} (J : Type*) [Category J] [Category D] + [HasColimitsOfShape J C] [HasColimitsOfShape J D] [HasExactColimitsOfShape J D] + (F : C ⥤ D) [PreservesFiniteLimits F] [ReflectsFiniteLimits F] [HasFiniteLimits C] + [PreservesColimitsOfShape J F] : HasExactColimitsOfShape J C where + preservesFiniteLimits := { preservesFiniteLimits I := { preservesLimit {G} := { + preserves {c} hc := by + constructor + apply isLimitOfReflects F + refine (IsLimit.equivOfNatIsoOfIso (isoWhiskerLeft G (preservesColimitNatIso F).symm) + ((_ ⋙ colim).mapCone c) _ ?_) (isLimitOfPreserves _ hc) + exact Cones.ext ((preservesColimitNatIso F).symm.app _) + fun i ↦ (preservesColimitNatIso F).inv.naturality _ } } } + +variable {C} in +/-- +Pull back a `HasExactLimitsOfShape J` along a functor which preserves and reflects finite colimits +and preserves limits of shape `J` +-/ +lemma HasExactLimitsOfShape.domain_of_functor {D : Type*} (J : Type*) [Category D] [Category J] + [HasLimitsOfShape J C] [HasLimitsOfShape J D] [HasExactLimitsOfShape J D] + (F : C ⥤ D) [PreservesFiniteColimits F] [ReflectsFiniteColimits F] [HasFiniteColimits C] + [PreservesLimitsOfShape J F] : HasExactLimitsOfShape J C where + preservesFiniteColimits := { preservesFiniteColimits I := { preservesColimit {G} := { + preserves {c} hc := by + constructor + apply isColimitOfReflects F + refine (IsColimit.equivOfNatIsoOfIso (isoWhiskerLeft G (preservesLimitNatIso F).symm) + ((_ ⋙ lim).mapCocone c) _ ?_) (isColimitOfPreserves _ hc) + refine Cocones.ext ((preservesLimitNatIso F).symm.app _) fun i ↦ ?_ + simp only [Functor.comp_obj, lim_obj, Functor.mapCocone_pt, isoWhiskerLeft_inv, Iso.symm_inv, + Cocones.precompose_obj_pt, whiskeringRight_obj_obj, Functor.const_obj_obj, + Cocones.precompose_obj_ι, NatTrans.comp_app, whiskerLeft_app, preservesLimitNatIso_hom_app, + Functor.mapCocone_ι_app, Functor.comp_map, whiskeringRight_obj_map, lim_map, Iso.app_hom, + Iso.symm_hom, preservesLimitNatIso_inv_app, Category.assoc] + rw [← Iso.eq_inv_comp] + exact (preservesLimitNatIso F).inv.naturality _ } } } + /-- Transport a `HasExactColimitsOfShape` along an equivalence of the shape. diff --git a/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms/FunctorCategory.lean b/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms/FunctorCategory.lean new file mode 100644 index 0000000000000..38ce82e5aefa3 --- /dev/null +++ b/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms/FunctorCategory.lean @@ -0,0 +1,29 @@ +/- +Copyright (c) 2024 Dagur Asgeirsson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Dagur Asgeirsson +-/ +import Mathlib.CategoryTheory.Abelian.GrothendieckAxioms.Basic +/-! + +# AB axioms in functor categories + +This file proves that, when the relevant limits and colimits exist, exactness of limits and +colimits carries over from `A` to the functor category `C ⥤ A` +-/ + +namespace CategoryTheory + +open CategoryTheory Limits Opposite + +variable {A C J : Type*} [Category A] [Category C] [Category J] + +instance [HasColimitsOfShape J A] [HasExactColimitsOfShape J A] [HasFiniteLimits A] : + HasExactColimitsOfShape J (C ⥤ A) where + preservesFiniteLimits := { preservesFiniteLimits _ := inferInstance } + +instance [HasLimitsOfShape J A] [HasExactLimitsOfShape J A] [HasFiniteColimits A] : + HasExactLimitsOfShape J (C ⥤ A) where + preservesFiniteColimits := { preservesFiniteColimits _ := inferInstance } + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms/Sheaf.lean b/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms/Sheaf.lean new file mode 100644 index 0000000000000..7696bed4e68af --- /dev/null +++ b/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms/Sheaf.lean @@ -0,0 +1,32 @@ +/- +Copyright (c) 2024 Dagur Asgeirsson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Dagur Asgeirsson +-/ +import Mathlib.CategoryTheory.Abelian.GrothendieckAxioms.FunctorCategory +import Mathlib.CategoryTheory.Sites.Limits +/-! + +# AB axioms in sheaf categories + +This file proves that, when the relevant limits and colimits and sheafification exist, exactness of +limits and colimits carries over from `A` to categories of `A`-valued sheaves. +-/ + +namespace CategoryTheory + +open Limits + +variable {A C J : Type*} [Category A] [Category C] [Category J] + +variable (K : GrothendieckTopology C) [HasWeakSheafify K A] + +instance [HasFiniteLimits A] [HasColimitsOfShape J A] [HasExactColimitsOfShape J A] + [PreservesColimitsOfShape J (sheafToPresheaf K A)] : HasExactColimitsOfShape J (Sheaf K A) := + HasExactColimitsOfShape.domain_of_functor J (sheafToPresheaf K A) + +instance [HasFiniteColimits A] [HasLimitsOfShape J A] [HasExactLimitsOfShape J A] + [PreservesFiniteColimits (sheafToPresheaf K A)] : HasExactLimitsOfShape J (Sheaf K A) := + HasExactLimitsOfShape.domain_of_functor J (sheafToPresheaf K A) + +end CategoryTheory diff --git a/Mathlib/Condensed/Light/AB.lean b/Mathlib/Condensed/Light/AB.lean index ce4794aa35c4e..fbafe40b55c8c 100644 --- a/Mathlib/Condensed/Light/AB.lean +++ b/Mathlib/Condensed/Light/AB.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Dagur Asgeirsson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Dagur Asgeirsson -/ -import Mathlib.CategoryTheory.Abelian.GrothendieckAxioms +import Mathlib.CategoryTheory.Abelian.GrothendieckAxioms.Basic import Mathlib.Condensed.Light.Epi import Mathlib.Condensed.Light.Limits /-! @@ -13,6 +13,7 @@ import Mathlib.Condensed.Light.Limits The category of light condensed `R`-modules over a ring satisfies the countable version of Grothendieck's AB4* axiom -/ + universe u open CategoryTheory Limits From df86f5952dd5c472884f826ceebec30eaa1a9e61 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Fri, 13 Dec 2024 10:51:11 +0000 Subject: [PATCH 696/829] feat(CategoryTheory/Sites): the property of being a sheaf of types is invariant under natural equivalences (#19921) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit If `P₁ : Cᵒᵖ ⥤ Type w` and `P₂ : Cᵒᵖ ⥤ Type w'` are naturally equivalent presheaves of types (in possibly different universes), then `P₁` is a sheaf for a Grothendieck topology iff `P₂` is. Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- Mathlib/CategoryTheory/Sites/IsSheafFor.lean | 46 +++++++++++++------ .../CategoryTheory/Sites/SheafOfTypes.lean | 26 ++++++++++- 2 files changed, 58 insertions(+), 14 deletions(-) diff --git a/Mathlib/CategoryTheory/Sites/IsSheafFor.lean b/Mathlib/CategoryTheory/Sites/IsSheafFor.lean index ce157ff0f2a75..45b5e94026382 100644 --- a/Mathlib/CategoryTheory/Sites/IsSheafFor.lean +++ b/Mathlib/CategoryTheory/Sites/IsSheafFor.lean @@ -65,7 +65,7 @@ which can be convenient. -/ -universe w v₁ v₂ u₁ u₂ +universe w w' v₁ v₂ u₁ u₂ namespace CategoryTheory @@ -603,21 +603,41 @@ theorem isSheafFor_top_sieve (P : Cᵒᵖ ⥤ Type w) : IsSheafFor P ((⊤ : Sie rw [← isSheafFor_iff_generate] apply isSheafFor_singleton_iso +/-- If `P₁ : Cᵒᵖ ⥤ Type w` and `P₂ : Cᵒᵖ ⥤ Type w` are two naturally equivalent +presheaves, and `P₁` is a sheaf for a presieve `R`, then `P₂` is also a sheaf for `R`. -/ +lemma isSheafFor_of_nat_equiv {P₁ : Cᵒᵖ ⥤ Type w} {P₂ : Cᵒᵖ ⥤ Type w'} + (e : ∀ ⦃X : C⦄, P₁.obj (op X) ≃ P₂.obj (op X)) + (he : ∀ ⦃X Y : C⦄ (f : X ⟶ Y) (x : P₁.obj (op Y)), + e (P₁.map f.op x) = P₂.map f.op (e x)) + {X : C} {R : Presieve X} (hP₁ : IsSheafFor P₁ R) : + IsSheafFor P₂ R := fun x₂ hx₂ ↦ by + have he' : ∀ ⦃X Y : C⦄ (f : X ⟶ Y) (x : P₂.obj (op Y)), + e.symm (P₂.map f.op x) = P₁.map f.op (e.symm x) := fun X Y f x ↦ + e.injective (by simp only [Equiv.apply_symm_apply, he]) + let x₁ : FamilyOfElements P₁ R := fun Y f hf ↦ e.symm (x₂ f hf) + have hx₁ : x₁.Compatible := fun Y₁ Y₂ Z g₁ g₂ f₁ f₂ h₁ h₂ fac ↦ e.injective + (by simp only [he, Equiv.apply_symm_apply, hx₂ g₁ g₂ h₁ h₂ fac, x₁]) + have : ∀ (t₂ : P₂.obj (op X)), + x₂.IsAmalgamation t₂ ↔ x₁.IsAmalgamation (e.symm t₂) := fun t₂ ↦ by + simp only [FamilyOfElements.IsAmalgamation, x₁, + ← he', EmbeddingLike.apply_eq_iff_eq] + refine ⟨e (hP₁.amalgamate x₁ hx₁), ?_, ?_⟩ + · dsimp + simp only [this, Equiv.symm_apply_apply] + exact IsSheafFor.isAmalgamation hP₁ hx₁ + · intro t₂ ht₂ + refine e.symm.injective ?_ + simp only [Equiv.symm_apply_apply] + exact hP₁.isSeparatedFor x₁ _ _ (by simpa only [this] using ht₂) + (IsSheafFor.isAmalgamation hP₁ hx₁) + /-- If `P` is a sheaf for `S`, and it is iso to `P'`, then `P'` is a sheaf for `S`. This shows that "being a sheaf for a presieve" is a mathematical or hygienic property. -/ -theorem isSheafFor_iso {P' : Cᵒᵖ ⥤ Type w} (i : P ≅ P') : IsSheafFor P R → IsSheafFor P' R := by - intro h x hx - let x' := x.compPresheafMap i.inv - have : x'.Compatible := FamilyOfElements.Compatible.compPresheafMap i.inv hx - obtain ⟨t, ht1, ht2⟩ := h x' this - use i.hom.app _ t - fconstructor - · convert FamilyOfElements.IsAmalgamation.compPresheafMap i.hom ht1 - simp [x'] - · intro y hy - rw [show y = (i.inv.app (op X) ≫ i.hom.app (op X)) y by simp] - simp [ht2 (i.inv.app _ y) (FamilyOfElements.IsAmalgamation.compPresheafMap i.inv hy)] +theorem isSheafFor_iso {P' : Cᵒᵖ ⥤ Type w} (i : P ≅ P') (hP : IsSheafFor P R) : + IsSheafFor P' R := + isSheafFor_of_nat_equiv (fun X ↦ (i.app (op X)).toEquiv) + (fun _ _ f x ↦ congr_fun (i.hom.naturality f.op) x) hP /-- If a presieve `R` on `X` has a subsieve `S` such that: diff --git a/Mathlib/CategoryTheory/Sites/SheafOfTypes.lean b/Mathlib/CategoryTheory/Sites/SheafOfTypes.lean index 9819869ea4a96..17f25b811da12 100644 --- a/Mathlib/CategoryTheory/Sites/SheafOfTypes.lean +++ b/Mathlib/CategoryTheory/Sites/SheafOfTypes.lean @@ -47,7 +47,7 @@ We also provide equivalent conditions to satisfy alternate definitions given in -/ -universe w v u +universe w w' v u namespace CategoryTheory @@ -82,6 +82,30 @@ theorem isSheaf_of_le (P : Cᵒᵖ ⥤ Type w) {J₁ J₂ : GrothendieckTopology theorem isSeparated_of_isSheaf (P : Cᵒᵖ ⥤ Type w) (h : IsSheaf J P) : IsSeparated J P := fun S hS => (h S hS).isSeparatedFor +section + +variable {J} {P₁ : Cᵒᵖ ⥤ Type w} {P₂ : Cᵒᵖ ⥤ Type w'} + (e : ∀ ⦃X : C⦄, P₁.obj (op X) ≃ P₂.obj (op X)) + (he : ∀ ⦃X Y : C⦄ (f : X ⟶ Y) (x : P₁.obj (op Y)), + e (P₁.map f.op x) = P₂.map f.op (e x)) + +include he in +lemma isSheaf_of_nat_equiv (hP₁ : Presieve.IsSheaf J P₁) : + Presieve.IsSheaf J P₂ := fun _ R hR ↦ + isSheafFor_of_nat_equiv e he (hP₁ R hR) + +include he in +lemma isSheaf_iff_of_nat_equiv : + Presieve.IsSheaf J P₁ ↔ Presieve.IsSheaf J P₂ := + ⟨fun hP₁ ↦ isSheaf_of_nat_equiv e he hP₁, + fun hP₂ ↦ + isSheaf_of_nat_equiv (fun _ ↦ (@e _).symm) (fun X Y f x ↦ by + obtain ⟨y, rfl⟩ := e.surjective x + refine e.injective ?_ + simp only [Equiv.apply_symm_apply, Equiv.symm_apply_apply, he]) hP₂⟩ + +end + /-- The property of being a sheaf is preserved by isomorphism. -/ theorem isSheaf_iso {P' : Cᵒᵖ ⥤ Type w} (i : P ≅ P') (h : IsSheaf J P) : IsSheaf J P' := fun _ S hS => isSheafFor_iso i (h S hS) From 0310194d4735fee2395a0b1fdbbaa29198a4b2a2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Iv=C3=A1n=20Renison?= <85908989+IvanRenison@users.noreply.github.com> Date: Fri, 13 Dec 2024 11:36:22 +0000 Subject: [PATCH 697/829] feat(Combinatorics/SimpleGraph): Add theorem `two_le_chromaticNumber_of_adj` (#19924) --- .../Combinatorics/SimpleGraph/ConcreteColorings.lean | 12 ++++++++++-- 1 file changed, 10 insertions(+), 2 deletions(-) diff --git a/Mathlib/Combinatorics/SimpleGraph/ConcreteColorings.lean b/Mathlib/Combinatorics/SimpleGraph/ConcreteColorings.lean index b968d04daf2bc..fdb7e6b18af60 100644 --- a/Mathlib/Combinatorics/SimpleGraph/ConcreteColorings.lean +++ b/Mathlib/Combinatorics/SimpleGraph/ConcreteColorings.lean @@ -19,6 +19,14 @@ This file defines colorings for some common graphs namespace SimpleGraph +theorem two_le_chromaticNumber_of_adj {α} {G : SimpleGraph α} {u v : α} (hAdj : G.Adj u v) : + 2 ≤ G.chromaticNumber := by + refine le_of_not_lt ?_ + intro h + have hc : G.Colorable 1 := chromaticNumber_le_iff_colorable.mp (Order.le_of_lt_add_one h) + let c : G.Coloring (Fin 1) := hc.some + exact c.valid hAdj (Subsingleton.elim (c u) (c v)) + /-- Bicoloring of a path graph -/ def pathGraph.bicoloring (n : ℕ) : Coloring (pathGraph n) Bool := @@ -43,8 +51,8 @@ theorem chromaticNumber_pathGraph (n : ℕ) (h : 2 ≤ n) : have hc := (pathGraph.bicoloring n).colorable apply le_antisymm · exact hc.chromaticNumber_le - · simpa only [pathGraph_two_eq_top, chromaticNumber_top] using - chromaticNumber_mono_of_embedding (pathGraph_two_embedding n h) + · have hAdj : (pathGraph n).Adj ⟨0, Nat.zero_lt_of_lt h⟩ ⟨1, h⟩ := by simp [pathGraph_adj] + exact two_le_chromaticNumber_of_adj hAdj theorem Coloring.even_length_iff_congr {α} {G : SimpleGraph α} (c : G.Coloring Bool) {u v : α} (p : G.Walk u v) : From 2f5c9675058cba44a8773c8e559b7d126b107570 Mon Sep 17 00:00:00 2001 From: Tomaz Gomes Mascarenhas Date: Fri, 13 Dec 2024 12:00:43 +0000 Subject: [PATCH 698/829] chore: shorten proof of le_iff_sub_nonneg (#19925) --- Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean | 4 +--- 1 file changed, 1 insertion(+), 3 deletions(-) diff --git a/Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean b/Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean index bb779638b5f04..21b69f6811d62 100644 --- a/Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean +++ b/Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean @@ -105,13 +105,11 @@ protected theorem le_iff_sub_nonneg (a b : ℚ) : a ≤ b ↔ 0 ≤ b - a := · apply Int.mul_nonneg h.2 (Int.natCast_nonneg _) · simp only [Int.natCast_pos, Nat.pos_iff_ne_zero] exact Nat.gcd_ne_zero_right (Nat.mul_ne_zero hb ha) - · simp only [divInt_ofNat, ← zero_iff_num_zero, mkRat_eq_zero hb] at h' - simp [h'] + · simp [h'] · simp only [Rat.sub_def, normalize_eq] refine ⟨fun H => ?_, fun H _ => ?_⟩ · refine Int.ediv_nonneg ?_ (Int.natCast_nonneg _) rw [Int.sub_nonneg] - push_neg at h obtain hb|hb := Ne.lt_or_lt h' · apply H intro H' From 658f6495b5cc142add209df9c45b314281a67d39 Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Fri, 13 Dec 2024 12:32:53 +0000 Subject: [PATCH 699/829] chore: move `Polynomial.basisMonomials` to a separate file (#19933) ... and generalize it to `Semiring`. Also add `Polynomial.toFinsuppIsoLinear` which is even true for `Semiring` and is in the very definition of `Polynomial` file. --- Mathlib.lean | 1 + Mathlib/Algebra/Polynomial/Basic.lean | 7 +++++ Mathlib/Algebra/Polynomial/Basis.lean | 30 ++++++++++++++++++++++ Mathlib/RingTheory/MvPolynomial/Basic.lean | 16 ------------ 4 files changed, 38 insertions(+), 16 deletions(-) create mode 100644 Mathlib/Algebra/Polynomial/Basis.lean diff --git a/Mathlib.lean b/Mathlib.lean index c63dcf2117b60..ed5502de95197 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -787,6 +787,7 @@ import Mathlib.Algebra.Periodic import Mathlib.Algebra.Pointwise.Stabilizer import Mathlib.Algebra.Polynomial.AlgebraMap import Mathlib.Algebra.Polynomial.Basic +import Mathlib.Algebra.Polynomial.Basis import Mathlib.Algebra.Polynomial.BigOperators import Mathlib.Algebra.Polynomial.Bivariate import Mathlib.Algebra.Polynomial.CancelLeads diff --git a/Mathlib/Algebra/Polynomial/Basic.lean b/Mathlib/Algebra/Polynomial/Basic.lean index 93d9b0870d987..8a361b32630cd 100644 --- a/Mathlib/Algebra/Polynomial/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Basic.lean @@ -346,6 +346,13 @@ def toFinsuppIso : R[X] ≃+* R[ℕ] where instance [DecidableEq R] : DecidableEq R[X] := @Equiv.decidableEq R[X] _ (toFinsuppIso R).toEquiv (Finsupp.instDecidableEq) +/-- Linear isomorphism between `R[X]` and `R[ℕ]`. This is just an +implementation detail, but it can be useful to transfer results from `Finsupp` to polynomials. -/ +@[simps!] +def toFinsuppIsoLinear : R[X] ≃ₗ[R] R[ℕ] where + __ := toFinsuppIso R + map_smul' _ _ := rfl + end AddMonoidAlgebra theorem ofFinsupp_sum {ι : Type*} (s : Finset ι) (f : ι → R[ℕ]) : diff --git a/Mathlib/Algebra/Polynomial/Basis.lean b/Mathlib/Algebra/Polynomial/Basis.lean new file mode 100644 index 0000000000000..ec2962085a92c --- /dev/null +++ b/Mathlib/Algebra/Polynomial/Basis.lean @@ -0,0 +1,30 @@ +/- +Copyright (c) 2024 Jz Pan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jz Pan +-/ +import Mathlib.Algebra.Polynomial.Basic +import Mathlib.LinearAlgebra.Basis.Defs + +/-! + +# Basis of a polynomial ring + +-/ + +universe u + +variable (R : Type u) [Semiring R] + +namespace Polynomial + +/-- The monomials form a basis on `R[X]`. To get the rank of a polynomial ring, +use this and `Basis.mk_eq_rank`. -/ +def basisMonomials : Basis ℕ R R[X] := + Basis.ofRepr (toFinsuppIsoLinear R) + +@[simp] +theorem coe_basisMonomials : (basisMonomials R : ℕ → R[X]) = fun s => monomial s 1 := + funext fun _ => ofFinsupp_single _ _ + +end Polynomial diff --git a/Mathlib/RingTheory/MvPolynomial/Basic.lean b/Mathlib/RingTheory/MvPolynomial/Basic.lean index 3fc4ad5148f15..b977a39a6f29c 100644 --- a/Mathlib/RingTheory/MvPolynomial/Basic.lean +++ b/Mathlib/RingTheory/MvPolynomial/Basic.lean @@ -5,7 +5,6 @@ Authors: Johannes Hölzl -/ import Mathlib.Algebra.CharP.Defs import Mathlib.Algebra.MvPolynomial.Degrees -import Mathlib.Algebra.Polynomial.AlgebraMap import Mathlib.Data.Fintype.Pi import Mathlib.LinearAlgebra.Finsupp.VectorSpace import Mathlib.LinearAlgebra.FreeModule.Finite.Basic @@ -40,8 +39,6 @@ noncomputable section open Set LinearMap Submodule -open Polynomial - universe u v variable (σ : Type u) (R : Type v) [CommSemiring R] (p m : ℕ) @@ -185,16 +182,3 @@ instance : IsScalarTower R (MvPolynomial σ R) (MvPolynomial σ S) := end Algebra end MvPolynomial - --- this is here to avoid import cycle issues -namespace Polynomial - -/-- The monomials form a basis on `R[X]`. -/ -noncomputable def basisMonomials : Basis ℕ R R[X] := - Basis.ofRepr (toFinsuppIsoAlg R).toLinearEquiv - -@[simp] -theorem coe_basisMonomials : (basisMonomials R : ℕ → R[X]) = fun s => monomial s 1 := - funext fun _ => ofFinsupp_single _ _ - -end Polynomial From ddbd495169452e1e42e3aa20d73da46b689a6b42 Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Fri, 13 Dec 2024 13:44:58 +0000 Subject: [PATCH 700/829] refactor(Algebra/Category/Ring): make morphisms a structure (#19757) Following the pattern of #19065, we introduce a `Hom` structure for morphisms in `SemiRingCat`, `CommSemiRingCat`, `RingCat` and `CommRingCat` allowing for many `simp`lified proofs and reducing the need of `erw`. Besides the additional explicitness in form of `ofHom` and `Hom.hom`, the main regression is friction with the `FunLike` instance from the general `ConcreteCategory` API, causing the need for explicit `rw`s between `f.hom x` and `(forget RingCat).map f x`. This regression can be fixed by replacing the general `ConcreteCategory.instFunLike` by a `CoeFun` (or `FunLike`) instance that is part of the `ConcreteCategory` data. Co-authored-by: Anne Baanen Co-authored-by: Anne C.A. Baanen --- .../Algebra/Category/BialgebraCat/Basic.lean | 14 +- Mathlib/Algebra/Category/BoolRing.lean | 92 +- .../Category/HopfAlgebraCat/Basic.lean | 14 +- .../ModuleCat/Differentials/Basic.lean | 57 +- .../ModuleCat/Differentials/Presheaf.lean | 5 +- .../Algebra/Category/ModuleCat/Presheaf.lean | 25 +- .../ModuleCat/Presheaf/ChangeOfRings.lean | 12 +- .../Category/ModuleCat/Presheaf/Colimits.lean | 16 +- .../Category/ModuleCat/Presheaf/Limits.lean | 8 +- .../Category/ModuleCat/Presheaf/Monoidal.lean | 22 +- .../ModuleCat/Presheaf/Pushforward.lean | 10 +- .../Category/ModuleCat/Presheaf/Sheafify.lean | 28 +- .../ModuleCat/Sheaf/ChangeOfRings.lean | 3 +- .../Algebra/Category/Ring/Adjunctions.lean | 17 +- Mathlib/Algebra/Category/Ring/Basic.lean | 952 ++++++++---------- Mathlib/Algebra/Category/Ring/Colimits.lean | 76 +- .../Algebra/Category/Ring/Constructions.lean | 119 +-- Mathlib/Algebra/Category/Ring/Epi.lean | 17 +- .../Category/Ring/FilteredColimits.lean | 70 +- Mathlib/Algebra/Category/Ring/Instances.lean | 39 +- Mathlib/Algebra/Category/Ring/Limits.lean | 34 +- .../Algebra/Category/Ring/LinearAlgebra.lean | 2 +- .../Algebra/Category/Ring/Under/Basic.lean | 16 +- .../Algebra/Category/Ring/Under/Limits.lean | 10 +- Mathlib/Algebra/Ring/Hom/Defs.lean | 3 + Mathlib/AlgebraicGeometry/AffineScheme.lean | 88 +- Mathlib/AlgebraicGeometry/AffineSpace.lean | 86 +- Mathlib/AlgebraicGeometry/FunctionField.lean | 20 +- .../GammaSpecAdjunction.lean | 30 +- Mathlib/AlgebraicGeometry/Modules/Tilde.lean | 7 +- .../AlgebraicGeometry/Morphisms/Affine.lean | 4 +- .../Morphisms/AffineAnd.lean | 38 +- .../Morphisms/ClosedImmersion.lean | 10 +- .../Morphisms/Constructors.lean | 8 +- .../AlgebraicGeometry/Morphisms/Finite.lean | 5 +- .../Morphisms/FinitePresentation.lean | 2 +- .../Morphisms/FiniteType.lean | 2 +- .../AlgebraicGeometry/Morphisms/Integral.lean | 4 +- .../Morphisms/Preimmersion.lean | 7 +- .../Morphisms/QuasiCompact.lean | 28 +- .../Morphisms/QuasiSeparated.lean | 87 +- .../Morphisms/RingHomProperties.lean | 75 +- .../Morphisms/Separated.lean | 2 +- .../AlgebraicGeometry/Morphisms/Smooth.lean | 10 +- .../Morphisms/SurjectiveOnStalks.lean | 9 +- .../Morphisms/UniversallyClosed.lean | 4 +- Mathlib/AlgebraicGeometry/OpenImmersion.lean | 16 +- .../ProjectiveSpectrum/Basic.lean | 5 +- .../ProjectiveSpectrum/Proper.lean | 4 +- .../ProjectiveSpectrum/Scheme.lean | 50 +- .../ProjectiveSpectrum/StructureSheaf.lean | 33 +- Mathlib/AlgebraicGeometry/Properties.lean | 30 +- Mathlib/AlgebraicGeometry/ResidueField.lean | 19 +- Mathlib/AlgebraicGeometry/Restrict.lean | 8 +- Mathlib/AlgebraicGeometry/Scheme.lean | 32 +- Mathlib/AlgebraicGeometry/Spec.lean | 107 +- Mathlib/AlgebraicGeometry/SpreadingOut.lean | 29 +- Mathlib/AlgebraicGeometry/Stalk.lean | 21 +- Mathlib/AlgebraicGeometry/StructureSheaf.lean | 123 +-- .../AlgebraicGeometry/ValuativeCriterion.lean | 31 +- Mathlib/Geometry/Manifold/Sheaf/Smooth.lean | 11 +- Mathlib/Geometry/RingedSpace/Basic.lean | 67 +- .../RingedSpace/LocallyRingedSpace.lean | 27 +- .../LocallyRingedSpace/HasColimits.lean | 50 +- .../LocallyRingedSpace/ResidueField.lean | 23 +- .../Geometry/RingedSpace/OpenImmersion.lean | 3 +- Mathlib/RingTheory/RingHomProperties.lean | 28 +- Mathlib/Topology/Category/TopCommRingCat.lean | 3 +- Mathlib/Topology/Sheaves/CommRingCat.lean | 107 +- .../CategoryTheory/ConcreteCategory/Ring.lean | 53 + 70 files changed, 1630 insertions(+), 1437 deletions(-) create mode 100644 MathlibTest/CategoryTheory/ConcreteCategory/Ring.lean diff --git a/Mathlib/Algebra/Category/BialgebraCat/Basic.lean b/Mathlib/Algebra/Category/BialgebraCat/Basic.lean index 59fe5b8fefba2..3dcc1b3f1497e 100644 --- a/Mathlib/Algebra/Category/BialgebraCat/Basic.lean +++ b/Mathlib/Algebra/Category/BialgebraCat/Basic.lean @@ -24,10 +24,13 @@ universe v u variable (R : Type u) [CommRing R] /-- The category of `R`-bialgebras. -/ -structure BialgebraCat extends Bundled Ring.{v} where - [instBialgebra : Bialgebra R α] +structure BialgebraCat where + /-- The underlying type. -/ + carrier : Type v + [instRing : Ring carrier] + [instBialgebra : Bialgebra R carrier] -attribute [instance] BialgebraCat.instBialgebra +attribute [instance] BialgebraCat.instBialgebra BialgebraCat.instRing variable {R} @@ -36,7 +39,7 @@ namespace BialgebraCat open Bialgebra instance : CoeSort (BialgebraCat.{v} R) (Type v) := - ⟨(·.α)⟩ + ⟨(·.carrier)⟩ variable (R) @@ -44,8 +47,7 @@ variable (R) @[simps] def of (X : Type v) [Ring X] [Bialgebra R X] : BialgebraCat R where - α := X - instBialgebra := (inferInstance : Bialgebra R X) + carrier := X variable {R} diff --git a/Mathlib/Algebra/Category/BoolRing.lean b/Mathlib/Algebra/Category/BoolRing.lean index 396dd1283f25f..0a8e91f15a4fa 100644 --- a/Mathlib/Algebra/Category/BoolRing.lean +++ b/Mathlib/Algebra/Category/BoolRing.lean @@ -23,52 +23,70 @@ universe u open CategoryTheory Order /-- The category of Boolean rings. -/ -def BoolRing := - Bundled BooleanRing +structure BoolRing where + private mk :: + /-- The underlying type. -/ + carrier : Type u + [booleanRing : BooleanRing carrier] namespace BoolRing +initialize_simps_projections BoolRing (-booleanRing) + instance : CoeSort BoolRing Type* := - Bundled.coeSort + ⟨carrier⟩ + +attribute [coe] carrier -instance (X : BoolRing) : BooleanRing X := - X.str +attribute [instance] booleanRing /-- Construct a bundled `BoolRing` from a `BooleanRing`. -/ -def of (α : Type*) [BooleanRing α] : BoolRing := - Bundled.of α +abbrev of (α : Type*) [BooleanRing α] : BoolRing := + ⟨α⟩ -@[simp] theorem coe_of (α : Type*) [BooleanRing α] : ↥(of α) = α := rfl instance : Inhabited BoolRing := ⟨of PUnit⟩ -instance : BundledHom.ParentProjection @BooleanRing.toCommRing := - ⟨⟩ - --- Porting note: `deriving` `ConcreteCategory` failed, added it manually --- see https://github.com/leanprover-community/mathlib4/issues/5020 -deriving instance LargeCategory for BoolRing - -instance : ConcreteCategory BoolRing := by - dsimp [BoolRing] - infer_instance - --- Porting note: disabled `simps` --- Invalid simp lemma BoolRing.hasForgetToCommRing_forget₂_obj_str_add. --- The given definition is not a constructor application: --- inferInstance.1 --- @[simps] -instance hasForgetToCommRing : HasForget₂ BoolRing CommRingCat := - BundledHom.forget₂ _ _ +variable {R} in +/-- The type of morphisms in `BoolRing`. -/ +@[ext] +structure Hom (R S : BoolRing) where + private mk :: + /-- The underlying ring hom. -/ + hom : R →+* S + +instance : Category BoolRing where + Hom R S := Hom R S + id R := ⟨RingHom.id R⟩ + comp f g := ⟨g.hom.comp f.hom⟩ + +@[ext] +lemma hom_ext {R S : BoolRing} {f g : R ⟶ S} (hf : f.hom = g.hom) : f = g := + Hom.ext hf + +/-- Typecheck a `RingHom` as a morphism in `BoolRing`. -/ +abbrev ofHom {R S : Type u} [BooleanRing R] [BooleanRing S] (f : R →+* S) : of R ⟶ of S := + ⟨f⟩ + +instance : ConcreteCategory BoolRing where + forget := + { obj := fun R ↦ R + map := fun f ↦ f.hom } + forget_faithful := ⟨fun h ↦ by ext x; simpa using congrFun h x⟩ + +instance hasForgetToCommRing : HasForget₂ BoolRing CommRingCat where + forget₂ := + { obj := fun R ↦ CommRingCat.of R + map := fun f ↦ CommRingCat.ofHom f.hom } /-- Constructs an isomorphism of Boolean rings from a ring isomorphism between them. -/ @[simps] def Iso.mk {α β : BoolRing.{u}} (e : α ≃+* β) : α ≅ β where - hom := (e : RingHom _ _) - inv := (e.symm : RingHom _ _) + hom := ⟨e⟩ + inv := ⟨e.symm⟩ hom_inv_id := by ext; exact e.symm_apply_apply _ inv_hom_id := by ext; exact e.apply_symm_apply _ @@ -76,22 +94,26 @@ end BoolRing /-! ### Equivalence between `BoolAlg` and `BoolRing` -/ -@[simps] -instance BoolRing.hasForgetToBoolAlg : HasForget₂ BoolRing BoolAlg where - forget₂ := - { obj := fun X => BoolAlg.of (AsBoolAlg X) - map := fun {_ _} => RingHom.asBoolAlg } - -- Porting note: Added. somehow it does not find this instance. instance {X : BoolAlg} : BooleanAlgebra ↑(BddDistLat.toBddLat (X.toBddDistLat)).toLat := BoolAlg.instBooleanAlgebra _ +instance {R : Type u} [BooleanRing R] : + BooleanRing (BoolAlg.of (AsBoolAlg ↑R)).toBddDistLat.toBddLat.toLat := + inferInstanceAs <| BooleanRing R + +@[simps] +instance BoolRing.hasForgetToBoolAlg : HasForget₂ BoolRing BoolAlg where + forget₂ := + { obj := fun X ↦ BoolAlg.of (AsBoolAlg X) + map := fun {R S} f ↦ RingHom.asBoolAlg f.hom } + @[simps] instance BoolAlg.hasForgetToBoolRing : HasForget₂ BoolAlg BoolRing where forget₂ := { obj := fun X => BoolRing.of (AsBoolRing X) - map := fun {_ _} => BoundedLatticeHom.asBoolRing } + map := fun {_ _} f => BoolRing.ofHom <| BoundedLatticeHom.asBoolRing f } /-- The equivalence between Boolean rings and Boolean algebras. This is actually an isomorphism. -/ @[simps functor inverse] diff --git a/Mathlib/Algebra/Category/HopfAlgebraCat/Basic.lean b/Mathlib/Algebra/Category/HopfAlgebraCat/Basic.lean index f1a3dfa0e2a0d..c5b834197885e 100644 --- a/Mathlib/Algebra/Category/HopfAlgebraCat/Basic.lean +++ b/Mathlib/Algebra/Category/HopfAlgebraCat/Basic.lean @@ -23,10 +23,13 @@ universe v u variable (R : Type u) [CommRing R] /-- The category of `R`-Hopf algebras. -/ -structure HopfAlgebraCat extends Bundled Ring.{v} where - [instHopfAlgebra : HopfAlgebra R α] +structure HopfAlgebraCat where + /-- The underlying type. -/ + carrier : Type v + [instRing : Ring carrier] + [instHopfAlgebra : HopfAlgebra R carrier] -attribute [instance] HopfAlgebraCat.instHopfAlgebra +attribute [instance] HopfAlgebraCat.instHopfAlgebra HopfAlgebraCat.instRing variable {R} @@ -35,7 +38,7 @@ namespace HopfAlgebraCat open HopfAlgebra instance : CoeSort (HopfAlgebraCat.{v} R) (Type v) := - ⟨(·.α)⟩ + ⟨(·.carrier)⟩ variable (R) @@ -43,8 +46,7 @@ variable (R) @[simps] def of (X : Type v) [Ring X] [HopfAlgebra R X] : HopfAlgebraCat R where - α := X - instHopfAlgebra := (inferInstance : HopfAlgebra R X) + carrier := X variable {R} diff --git a/Mathlib/Algebra/Category/ModuleCat/Differentials/Basic.lean b/Mathlib/Algebra/Category/ModuleCat/Differentials/Basic.lean index a0ca4297cf81c..5e7c340747c0f 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Differentials/Basic.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Differentials/Basic.lean @@ -31,8 +31,8 @@ variable {A B : CommRingCat.{u}} (M : ModuleCat.{v} B) (f : A ⟶ B) /-- The type of derivations with values in a `B`-module `M` relative to a morphism `f : A ⟶ B` in the category `CommRingCat`. -/ nonrec def Derivation : Type _ := - letI := f.toAlgebra - letI := Module.compHom M f + letI := f.hom.toAlgebra + letI := Module.compHom M f.hom Derivation A B M namespace Derivation @@ -44,8 +44,8 @@ def mk (d : B → M) (d_add : ∀ (b b' : B), d (b + b') = d b + d b' := by simp (d_mul : ∀ (b b' : B), d (b * b') = b • d b' + b' • d b := by simp) (d_map : ∀ (a : A), d (f a) = 0 := by simp) : M.Derivation f := - letI := f.toAlgebra - letI := Module.compHom M f + letI := f.hom.toAlgebra + letI := Module.compHom M f.hom { toFun := d map_add' := d_add map_smul' := fun a b ↦ by @@ -54,7 +54,7 @@ def mk (d : B → M) (d_add : ∀ (b b' : B), d (b + b') = d b + d b' := by simp rfl map_one_eq_zero' := by dsimp - rw [← f.map_one, d_map] + rw [← f.hom.map_one, d_map] leibniz' := d_mul } variable (D : M.Derivation f) @@ -62,8 +62,8 @@ variable (D : M.Derivation f) /-- The underlying map `B → M` of a derivation `M.Derivation f` when `M : ModuleCat B` and `f : A ⟶ B` is a morphism in `CommRingCat`. -/ def d (b : B) : M := - letI := f.toAlgebra - letI := Module.compHom M f + letI := f.hom.toAlgebra + letI := Module.compHom M f.hom _root_.Derivation.toLinearMap D b @[simp] @@ -74,8 +74,8 @@ lemma d_mul (b b' : B) : D.d (b * b') = b • D.d b' + b' • D.d b := by simp [ @[simp] lemma d_map (a : A) : D.d (f a) = 0 := - letI := f.toAlgebra - letI := Module.compHom M f + letI := f.hom.toAlgebra + letI := Module.compHom M f.hom D.map_algebraMap a end Derivation @@ -90,7 +90,7 @@ variable {A B A' B' : CommRingCat.{u}} {f : A ⟶ B} {f' : A' ⟶ B'} variable (f) in /-- The module of differentials of a morphism `f : A ⟶ B` in the category `CommRingCat`. -/ noncomputable def KaehlerDifferential : ModuleCat.{u} B := - letI := f.toAlgebra + letI := f.hom.toAlgebra ModuleCat.of B (_root_.KaehlerDifferential A B) namespace KaehlerDifferential @@ -99,9 +99,9 @@ variable (f) in /-- The (universal) derivation in `(KaehlerDifferential f).Derivation f` when `f : A ⟶ B` is a morphism in the category `CommRingCat`. -/ noncomputable def D : (KaehlerDifferential f).Derivation f := - letI algebra := f.toAlgebra + letI := f.hom.toAlgebra ModuleCat.Derivation.mk - (fun b ↦ _root_.KaehlerDifferential.D A B b) (by simp [algebra]) (by simp [algebra]) + (fun b ↦ _root_.KaehlerDifferential.D A B b) (by simp) (by simp) (_root_.KaehlerDifferential.D A B).map_algebraMap /-- When `f : A ⟶ B` is a morphism in the category `CommRingCat`, this is the @@ -126,30 +126,25 @@ induced by a commutative square (given by an equality `g ≫ f' = f ≫ g'`) in the category `CommRingCat`. -/ noncomputable def map : KaehlerDifferential f ⟶ - (ModuleCat.restrictScalars g').obj (KaehlerDifferential f') := - letI := f.toAlgebra - letI := f'.toAlgebra - letI := g.toAlgebra - letI := g'.toAlgebra - letI := (g ≫ f').toAlgebra + (ModuleCat.restrictScalars g'.hom).obj (KaehlerDifferential f') := + letI := f.hom.toAlgebra + letI := f'.hom.toAlgebra + letI := g.hom.toAlgebra + letI := g'.hom.toAlgebra + letI := (g ≫ f').hom.toAlgebra have : IsScalarTower A A' B' := IsScalarTower.of_algebraMap_eq' rfl - have := IsScalarTower.of_algebraMap_eq' fac + have := IsScalarTower.of_algebraMap_eq' (congrArg Hom.hom fac) -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(Y := ...)`. -- This suggests `restrictScalars` needs to be redesigned. - ModuleCat.ofHom (Y := (ModuleCat.restrictScalars g').obj (KaehlerDifferential f')) + ModuleCat.ofHom (Y := (ModuleCat.restrictScalars g'.hom).obj (KaehlerDifferential f')) { toFun := fun x ↦ _root_.KaehlerDifferential.map A A' B B' x map_add' := by simp map_smul' := by simp } @[simp] lemma map_d (b : B) : map fac (d b) = d (g' b) := by - letI := f.toAlgebra - letI := f'.toAlgebra - letI := g.toAlgebra - letI := g'.toAlgebra - letI := (f'.comp g).toAlgebra - have : IsScalarTower A A' B' := IsScalarTower.of_algebraMap_eq' rfl - have := IsScalarTower.of_algebraMap_eq' fac + algebraize [f.hom, f'.hom, g.hom, g'.hom, f'.hom.comp g.hom] + have := IsScalarTower.of_algebraMap_eq' (congrArg Hom.hom fac) exact _root_.KaehlerDifferential.map_D A A' B B' b end KaehlerDifferential @@ -165,14 +160,14 @@ variable {A B : CommRingCat.{u}} {f : A ⟶ B} and `D : M.Derivation f`, this is the induced morphism `CommRingCat.KaehlerDifferential f ⟶ M`. -/ noncomputable def desc : CommRingCat.KaehlerDifferential f ⟶ M := - letI := f.toAlgebra - letI := Module.compHom M f + letI := f.hom.toAlgebra + letI := Module.compHom M f.hom ofHom D.liftKaehlerDifferential @[simp] lemma desc_d (b : B) : D.desc (CommRingCat.KaehlerDifferential.d b) = D.d b := by - letI := f.toAlgebra - letI := Module.compHom M f + letI := f.hom.toAlgebra + letI := Module.compHom M f.hom apply D.liftKaehlerDifferential_comp_D end ModuleCat.Derivation diff --git a/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean b/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean index 9df2e8a46993e..da141704256b2 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean @@ -182,8 +182,9 @@ noncomputable def relativeDifferentials' : PresheafOfModules.{u} (R ⋙ forget₂ _ _) where obj X := CommRingCat.KaehlerDifferential (φ'.app X) map f := CommRingCat.KaehlerDifferential.map (φ'.naturality f) - map_id _ := by ext; simp; rfl - map_comp _ _ := by ext; simp + -- Without `dsimp`, `ext` doesn't pick up the right lemmas. + map_id _ := by dsimp; ext; simp + map_comp _ _ := by dsimp; ext; simp attribute [simp] relativeDifferentials'_obj diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean index 8755d8d2ac2f7..a619b7e79f880 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean @@ -40,12 +40,15 @@ structure PresheafOfModules where /-- a family of modules over `R.obj X` for all `X` -/ obj (X : Cᵒᵖ) : ModuleCat.{v} (R.obj X) /-- the restriction maps of a presheaf of modules -/ - map {X Y : Cᵒᵖ} (f : X ⟶ Y) : obj X ⟶ (ModuleCat.restrictScalars (R.map f)).obj (obj Y) + map {X Y : Cᵒᵖ} (f : X ⟶ Y) : obj X ⟶ (ModuleCat.restrictScalars (R.map f).hom).obj (obj Y) map_id (X : Cᵒᵖ) : - map (𝟙 X) = (ModuleCat.restrictScalarsId' _ (R.map_id X)).inv.app _ := by aesop_cat + map (𝟙 X) = + (ModuleCat.restrictScalarsId' _ (congrArg RingCat.Hom.hom (R.map_id X))).inv.app _ := by + aesop_cat map_comp {X Y Z : Cᵒᵖ} (f : X ⟶ Y) (g : Y ⟶ Z) : map (f ≫ g) = map f ≫ (ModuleCat.restrictScalars _).map (map g) ≫ - (ModuleCat.restrictScalarsComp' _ _ _ (R.map_comp f g)).inv.app _ := by aesop_cat + (ModuleCat.restrictScalarsComp' _ _ _ + (congrArg RingCat.Hom.hom <| R.map_comp f g)).inv.app _ := by aesop_cat namespace PresheafOfModules @@ -67,7 +70,7 @@ structure Hom where /-- a family of linear maps `M₁.obj X ⟶ M₂.obj X` for all `X`. -/ app (X : Cᵒᵖ) : M₁.obj X ⟶ M₂.obj X naturality {X Y : Cᵒᵖ} (f : X ⟶ Y) : - M₁.map f ≫ (ModuleCat.restrictScalars (R.map f)).map (app Y) = + M₁.map f ≫ (ModuleCat.restrictScalars (R.map f).hom).map (app Y) = app X ≫ M₂.map f := by aesop_cat attribute [reassoc (attr := simp)] Hom.naturality @@ -100,7 +103,7 @@ lemma naturality_apply (f : M₁ ⟶ M₂) {X Y : Cᵒᵖ} (g : X ⟶ Y) (x : M @[simps!] def isoMk (app : ∀ (X : Cᵒᵖ), M₁.obj X ≅ M₂.obj X) (naturality : ∀ ⦃X Y : Cᵒᵖ⦄ (f : X ⟶ Y), - M₁.map f ≫ (ModuleCat.restrictScalars (R.map f)).map (app Y).hom = + M₁.map f ≫ (ModuleCat.restrictScalars (R.map f).hom).map (app Y).hom = (app X).hom ≫ M₂.map f := by aesop_cat) : M₁ ≅ M₂ where hom := { app := fun X ↦ (app X).hom } inv := @@ -164,7 +167,7 @@ def ofPresheaf : PresheafOfModules.{v} R where -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(Y := ...)`. -- This suggests `restrictScalars` needs to be redesigned. map {X Y} f := ModuleCat.ofHom - (Y := (ModuleCat.restrictScalars (R.map f)).obj (ModuleCat.of _ (M.obj Y))) + (Y := (ModuleCat.restrictScalars (R.map f).hom).obj (ModuleCat.of _ (M.obj Y))) { toFun := fun x ↦ M.map f x map_add' := by simp map_smul' := fun r m ↦ map_smul f r m } @@ -255,7 +258,7 @@ instance (X : Cᵒᵖ) : (evaluation.{v} R X).Additive where to restriction of scalars. -/ @[simps] noncomputable def restriction {X Y : Cᵒᵖ} (f : X ⟶ Y) : - evaluation R X ⟶ evaluation R Y ⋙ ModuleCat.restrictScalars (R.map f) where + evaluation R X ⟶ evaluation R Y ⋙ ModuleCat.restrictScalars (R.map f).hom where app M := M.map f /-- The obvious free presheaf of modules of rank `1`. -/ @@ -264,13 +267,13 @@ def unit : PresheafOfModules R where -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(Y := ...)`. -- This suggests `restrictScalars` needs to be redesigned. map {X Y} f := ModuleCat.ofHom - (Y := (ModuleCat.restrictScalars (R.map f)).obj (ModuleCat.of (R.obj Y) (R.obj Y))) + (Y := (ModuleCat.restrictScalars (R.map f).hom).obj (ModuleCat.of (R.obj Y) (R.obj Y))) { toFun := fun x ↦ R.map f x map_add' := by simp map_smul' := by aesop_cat } lemma unit_map_one {X Y : Cᵒᵖ} (f : X ⟶ Y) : (unit R).map f (1 : R.obj X) = (1 : R.obj Y) := - (R.map f).map_one + (R.map f).hom.map_one variable {R} @@ -349,7 +352,7 @@ variable (M : PresheafOfModules.{v} R) /-- Auxiliary definition for `forgetToPresheafModuleCatObj`. -/ noncomputable abbrev forgetToPresheafModuleCatObjObj (Y : Cᵒᵖ) : ModuleCat (R.obj X) := - (ModuleCat.restrictScalars (R.map (hX.to Y))).obj (M.obj Y) + (ModuleCat.restrictScalars (R.map (hX.to Y)).hom).obj (M.obj Y) -- This should not be a `simp` lemma because `M.obj Y` is missing the `Module (R.obj X)` instance, -- so `simp`ing breaks downstream proofs. @@ -367,7 +370,7 @@ def forgetToPresheafModuleCatObjMap {Y Z : Cᵒᵖ} (f : Y ⟶ Z) : map_smul' := fun r x => by simp only [ModuleCat.restrictScalars.smul_def, AddHom.toFun_eq_coe, AddHom.coe_mk, RingHom.id_apply, M.map_smul] - rw [← CategoryTheory.comp_apply, ← R.map_comp] + rw [← RingCat.comp_apply, ← R.map_comp] congr apply hX.hom_ext } diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/ChangeOfRings.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/ChangeOfRings.lean index 2b6d781c0892b..57574ede2b800 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/ChangeOfRings.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/ChangeOfRings.lean @@ -27,17 +27,17 @@ variable {C : Type u'} [Category.{v'} C] {R R' : Cᵒᵖ ⥤ RingCat.{u}} @[simps] noncomputable def restrictScalarsObj (M' : PresheafOfModules.{v} R') (α : R ⟶ R') : PresheafOfModules R where - obj := fun X ↦ (ModuleCat.restrictScalars (α.app X)).obj (M'.obj X) + obj := fun X ↦ (ModuleCat.restrictScalars (α.app X).hom).obj (M'.obj X) -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(X := ...)` and `(Y := ...)`. -- This suggests `restrictScalars` needs to be redesigned. map := fun {X Y} f ↦ ModuleCat.ofHom - (X := (ModuleCat.restrictScalars (α.app X)).obj (M'.obj X)) - (Y := (ModuleCat.restrictScalars (R.map f)).obj - ((ModuleCat.restrictScalars (α.app Y)).obj (M'.obj Y))) + (X := (ModuleCat.restrictScalars (α.app X).hom).obj (M'.obj X)) + (Y := (ModuleCat.restrictScalars (R.map f).hom).obj + ((ModuleCat.restrictScalars (α.app Y).hom).obj (M'.obj Y))) { toFun := M'.map f map_add' := map_add _ map_smul' := fun r x ↦ (M'.map_smul f (α.app _ r) x).trans (by - have eq := RingHom.congr_fun (α.naturality f) r + have eq := RingHom.congr_fun (congrArg RingCat.Hom.hom <| α.naturality f) r dsimp at eq rw [← eq] rfl ) } @@ -49,7 +49,7 @@ noncomputable def restrictScalars (α : R ⟶ R') : PresheafOfModules.{v} R' ⥤ PresheafOfModules.{v} R where obj M' := M'.restrictScalarsObj α map φ' := - { app := fun X ↦ (ModuleCat.restrictScalars (α.app X)).map (Hom.app φ' X) + { app := fun X ↦ (ModuleCat.restrictScalars (α.app X).hom).map (Hom.app φ' X) naturality := fun {X Y} f ↦ by ext x exact naturality_apply φ' f x } diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Colimits.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Colimits.lean index fa22a86db7365..d5a162c23a5b5 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Colimits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Colimits.lean @@ -26,7 +26,7 @@ variable {C : Type u₁} [Category.{v₁} C] {R : Cᵒᵖ ⥤ RingCat.{u}} section Colimits variable [∀ {X Y : Cᵒᵖ} (f : X ⟶ Y), PreservesColimit (F ⋙ evaluation R Y) - (ModuleCat.restrictScalars (R.map f))] + (ModuleCat.restrictScalars (R.map f).hom)] /-- A cocone in the category `PresheafOfModules R` is colimit if it is so after the application of the functors `evaluation R X` for all `X`. -/ @@ -54,8 +54,8 @@ def evaluationJointlyReflectsColimits (c : Cocone F) variable [∀ X, HasColimit (F ⋙ evaluation R X)] instance {X Y : Cᵒᵖ} (f : X ⟶ Y) : - HasColimit (F ⋙ evaluation R Y ⋙ (ModuleCat.restrictScalars (R.map f))) := - ⟨_, isColimitOfPreserves (ModuleCat.restrictScalars (R.map f)) + HasColimit (F ⋙ evaluation R Y ⋙ (ModuleCat.restrictScalars (R.map f).hom)) := + ⟨_, isColimitOfPreserves (ModuleCat.restrictScalars (R.map f).hom) (colimit.isColimit (F ⋙ evaluation R Y))⟩ /-- Given `F : J ⥤ PresheafOfModules.{v} R`, this is the presheaf of modules obtained by @@ -64,21 +64,21 @@ taking a colimit in the category of modules over `R.obj X` for all `X`. -/ noncomputable def colimitPresheafOfModules : PresheafOfModules R where obj X := colimit (F ⋙ evaluation R X) map {_ Y} f := colimMap (whiskerLeft F (restriction R f)) ≫ - (preservesColimitIso (ModuleCat.restrictScalars (R.map f)) (F ⋙ evaluation R Y)).inv + (preservesColimitIso (ModuleCat.restrictScalars (R.map f).hom) (F ⋙ evaluation R Y)).inv map_id X := colimit.hom_ext (fun j => by dsimp rw [ι_colimMap_assoc, whiskerLeft_app, restriction_app] - erw [ι_preservesColimitIso_inv (G := ModuleCat.restrictScalars (R.map (𝟙 X))), + erw [ι_preservesColimitIso_inv (G := ModuleCat.restrictScalars (R.map (𝟙 X)).hom), ModuleCat.restrictScalarsId'App_inv_naturality] rw [map_id] dsimp) map_comp {X Y Z} f g := colimit.hom_ext (fun j => by dsimp rw [ι_colimMap_assoc, whiskerLeft_app, restriction_app, assoc, ι_colimMap_assoc] - erw [ι_preservesColimitIso_inv (G := ModuleCat.restrictScalars (R.map (f ≫ g))), - ι_preservesColimitIso_inv_assoc (G := ModuleCat.restrictScalars (R.map f))] + erw [ι_preservesColimitIso_inv (G := ModuleCat.restrictScalars (R.map (f ≫ g)).hom), + ι_preservesColimitIso_inv_assoc (G := ModuleCat.restrictScalars (R.map f).hom)] rw [← Functor.map_comp_assoc, ι_colimMap_assoc] - erw [ι_preservesColimitIso_inv (G := ModuleCat.restrictScalars (R.map g))] + erw [ι_preservesColimitIso_inv (G := ModuleCat.restrictScalars (R.map g).hom)] rw [map_comp, ModuleCat.restrictScalarsComp'_inv_app, assoc, assoc, whiskerLeft_app, whiskerLeft_app, restriction_app, restriction_app] simp only [Functor.map_comp, assoc] diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Limits.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Limits.lean index 12b4f9fc866eb..e6ecca097ac2a 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Limits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Limits.lean @@ -36,7 +36,7 @@ def evaluationJointlyReflectsLimits (c : Cone F) lift s := { app := fun X => (hc X).lift ((evaluation R X).mapCone s) naturality := fun {X Y} f ↦ by - apply (isLimitOfPreserves (ModuleCat.restrictScalars (R.map f)) (hc Y)).hom_ext + apply (isLimitOfPreserves (ModuleCat.restrictScalars (R.map f).hom) (hc Y)).hom_ext intro j have h₁ := (c.π.app j).naturality f have h₂ := (hc X).fac ((evaluation R X).mapCone s) j @@ -54,8 +54,8 @@ def evaluationJointlyReflectsLimits (c : Cone F) rw [← hm, comp_app] instance {X Y : Cᵒᵖ} (f : X ⟶ Y) : - HasLimit (F ⋙ evaluation R Y ⋙ ModuleCat.restrictScalars (R.map f)) := by - change HasLimit ((F ⋙ evaluation R Y) ⋙ ModuleCat.restrictScalars (R.map f)) + HasLimit (F ⋙ evaluation R Y ⋙ ModuleCat.restrictScalars (R.map f).hom) := by + change HasLimit ((F ⋙ evaluation R Y) ⋙ ModuleCat.restrictScalars (R.map f).hom) infer_instance /-- Given `F : J ⥤ PresheafOfModules.{v} R`, this is the presheaf of modules obtained by @@ -64,7 +64,7 @@ taking a limit in the category of modules over `R.obj X` for all `X`. -/ noncomputable def limitPresheafOfModules : PresheafOfModules R where obj X := limit (F ⋙ evaluation R X) map {_ Y} f := limMap (whiskerLeft F (restriction R f)) ≫ - (preservesLimitIso (ModuleCat.restrictScalars (R.map f)) (F ⋙ evaluation R Y)).inv + (preservesLimitIso (ModuleCat.restrictScalars (R.map f).hom) (F ⋙ evaluation R Y)).inv map_id X := by dsimp rw [← cancel_mono (preservesLimitIso _ _).hom, assoc, Iso.inv_hom_id, comp_id] diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean index f4d23fbb2a3e3..9547c9239b0fc 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean @@ -38,7 +38,7 @@ variable (M₁ M₂ M₃ M₄ : PresheafOfModules.{u} (R ⋙ forget₂ _ _)) /-- Auxiliary definition for `tensorObj`. -/ noncomputable def tensorObjMap {X Y : Cᵒᵖ} (f : X ⟶ Y) : M₁.obj X ⊗ M₂.obj X ⟶ - (ModuleCat.restrictScalars (R.map f)).obj (M₁.obj Y ⊗ M₂.obj Y) := + (ModuleCat.restrictScalars (R.map f).hom).obj (M₁.obj Y ⊗ M₂.obj Y) := ModuleCat.MonoidalCategory.tensorLift (fun m₁ m₂ ↦ M₁.map f m₁ ⊗ₜ M₂.map f m₂) (by intro m₁ m₁' m₂; dsimp; rw [map_add, TensorProduct.add_tmul]) (by intro a m₁ m₂; dsimp; erw [M₁.map_smul]; rfl) @@ -53,24 +53,18 @@ noncomputable def tensorObj : PresheafOfModules (R ⋙ forget₂ _ _) where map_id X := ModuleCat.MonoidalCategory.tensor_ext (by intro m₁ m₂ dsimp [tensorObjMap] - simp only [map_id, Functor.comp_obj, CommRingCat.forgetToRingCat_obj, Functor.comp_map, - ModuleCat.restrictScalarsId'_inv_app, ModuleCat.restrictScalarsId'App_inv_apply] - rfl) + simp) map_comp f g := ModuleCat.MonoidalCategory.tensor_ext (by intro m₁ m₂ dsimp [tensorObjMap] - simp only [map_comp, Functor.comp_obj, CommRingCat.forgetToRingCat_obj, Functor.comp_map, - ModuleCat.restrictScalarsComp'_inv_app, ModuleCat.hom_comp, LinearMap.comp_apply, - ModuleCat.restrictScalars.map_apply, ModuleCat.restrictScalarsComp'App_inv_apply] - rfl) + simp) variable {M₁ M₂ M₃ M₄} @[simp] lemma tensorObj_map_tmul {X Y : Cᵒᵖ} (f : X ⟶ Y) (m₁ : M₁.obj X) (m₂ : M₂.obj X) : DFunLike.coe (α := (M₁.obj X ⊗ M₂.obj X : _)) - (β := fun _ ↦ (ModuleCat.restrictScalars - ((forget₂ CommRingCat RingCat).map (R.map f))).obj (M₁.obj Y ⊗ M₂.obj Y)) + (β := fun _ ↦ (ModuleCat.restrictScalars (R.map f).hom).obj (M₁.obj Y ⊗ M₂.obj Y)) ((tensorObj M₁ M₂).map f).hom (m₁ ⊗ₜ[R.obj X] m₂) = M₁.map f m₁ ⊗ₜ[R.obj Y] M₂.map f m₂ := rfl /-- The tensor product of two morphisms of presheaves of modules. -/ @@ -98,13 +92,13 @@ noncomputable instance monoidalCategoryStruct : (fun _ _ _ ↦ ModuleCat.MonoidalCategory.tensor_ext₃' (by intros; rfl)) leftUnitor M := Iso.symm (isoMk (fun _ ↦ (λ_ _).symm) (fun X Y f ↦ by ext m - dsimp - erw [leftUnitor_inv_apply, leftUnitor_inv_apply, tensorObj_map_tmul, (R.map f).map_one] + dsimp [CommRingCat.forgetToRingCat_obj] + erw [leftUnitor_inv_apply, leftUnitor_inv_apply, tensorObj_map_tmul, (R.map f).hom.map_one] rfl)) rightUnitor M := Iso.symm (isoMk (fun _ ↦ (ρ_ _).symm) (fun X Y f ↦ by ext m - dsimp - erw [rightUnitor_inv_apply, rightUnitor_inv_apply, tensorObj_map_tmul, (R.map f).map_one] + dsimp [CommRingCat.forgetToRingCat_obj] + erw [rightUnitor_inv_apply, rightUnitor_inv_apply, tensorObj_map_tmul, (R.map f).hom.map_one] rfl)) noncomputable instance monoidalCategory : diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Pushforward.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Pushforward.lean index 099ab943de7ac..62c75c356bf58 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Pushforward.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Pushforward.lean @@ -71,26 +71,26 @@ noncomputable def pushforwardCompToPresheaf : Iso.refl _ lemma pushforward_obj_map_apply (M : PresheafOfModules.{v} R) {X Y : Cᵒᵖ} (f : X ⟶ Y) - (m : (ModuleCat.restrictScalars (φ.app X)).obj (M.obj (Opposite.op (F.obj X.unop)))) : + (m : (ModuleCat.restrictScalars (φ.app X).hom).obj (M.obj (Opposite.op (F.obj X.unop)))) : (((pushforward φ).obj M).map f).hom m = M.map (F.map f.unop).op m := rfl /-- `@[simp]`-normal form of `pushforward_obj_map_apply`. -/ @[simp] lemma pushforward_obj_map_apply' (M : PresheafOfModules.{v} R) {X Y : Cᵒᵖ} (f : X ⟶ Y) - (m : (ModuleCat.restrictScalars (φ.app X)).obj (M.obj (Opposite.op (F.obj X.unop)))) : + (m : (ModuleCat.restrictScalars (φ.app X).hom).obj (M.obj (Opposite.op (F.obj X.unop)))) : DFunLike.coe (F := ↑((ModuleCat.restrictScalars _).obj _) →ₗ[_] - ↑((ModuleCat.restrictScalars (S.map f)).obj ((ModuleCat.restrictScalars _).obj _))) + ↑((ModuleCat.restrictScalars (S.map f).hom).obj ((ModuleCat.restrictScalars _).obj _))) (((pushforward φ).obj M).map f).hom m = M.map (F.map f.unop).op m := rfl lemma pushforward_map_app_apply {M N : PresheafOfModules.{v} R} (α : M ⟶ N) (X : Cᵒᵖ) - (m : (ModuleCat.restrictScalars (φ.app X)).obj (M.obj (Opposite.op (F.obj X.unop)))) : + (m : (ModuleCat.restrictScalars (φ.app X).hom).obj (M.obj (Opposite.op (F.obj X.unop)))) : (((pushforward φ).map α).app X).hom m = α.app (Opposite.op (F.obj X.unop)) m := rfl /-- `@[simp]`-normal form of `pushforward_map_app_apply`. -/ @[simp] lemma pushforward_map_app_apply' {M N : PresheafOfModules.{v} R} (α : M ⟶ N) (X : Cᵒᵖ) - (m : (ModuleCat.restrictScalars (φ.app X)).obj (M.obj (Opposite.op (F.obj X.unop)))) : + (m : (ModuleCat.restrictScalars (φ.app X).hom).obj (M.obj (Opposite.op (F.obj X.unop)))) : DFunLike.coe (F := ↑((ModuleCat.restrictScalars _).obj _) →ₗ[_] ↑((ModuleCat.restrictScalars _).obj _)) (((pushforward φ).map α).app X).hom m = α.app (Opposite.op (F.obj X.unop)) m := rfl diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean index bcf7dd4cfc6a3..d1dec86a2c040 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean @@ -61,7 +61,8 @@ lemma _root_.PresheafOfModules.Sheafify.app_eq_of_isLocallyInjective (hr₀ : α.app _ r₀ = α.app _ r₀') (hm₀ : φ.app _ m₀ = φ.app _ m₀') : φ.app _ (r₀ • m₀) = φ.app _ (r₀' • m₀') := by - apply hA _ (Presheaf.equalizerSieve r₀ r₀' ⊓ Presheaf.equalizerSieve (F := M₀.presheaf) m₀ m₀') + apply hA _ (Presheaf.equalizerSieve (D := RingCat) r₀ r₀' ⊓ + Presheaf.equalizerSieve (F := M₀.presheaf) m₀ m₀') · apply J.intersection_covering · exact Presheaf.equalizerSieve_mem J α _ _ hr₀ · exact Presheaf.equalizerSieve_mem J φ _ _ hm₀ @@ -78,7 +79,8 @@ lemma isCompatible_map_smul_aux {Y Z : C} (f : Y ⟶ X) (g : Z ⟶ Y) φ.app _ (M₀.map g.op (r₀ • m₀)) = φ.app _ (r₀' • m₀') := by rw [← PresheafOfModules.Sheafify.app_eq_of_isLocallyInjective α φ hA (R₀.map g.op r₀) r₀' (M₀.map g.op m₀) m₀', M₀.map_smul] - · rw [hr₀', R.map_comp, comp_apply, ← hr₀, NatTrans.naturality_apply] + · rw [hr₀', R.map_comp, RingCat.comp_apply, ← hr₀, ← RingCat.comp_apply, NatTrans.naturality, + RingCat.comp_apply] · rw [hm₀', A.map_comp, AddCommGrp.coe_comp, Function.comp_apply, ← hm₀] erw [NatTrans.naturality_apply] @@ -100,7 +102,8 @@ lemma isCompatible_map_smul : ((r₀.smul m₀).map (whiskerRight φ (forget _)) have hb₂ : (φ.app (Opposite.op Y₂)) b₂ = (A.map f₂.op) m := (hm₀ f₂ h₂).symm have ha₀ : (α.app (Opposite.op Z)) a₀ = (R.map (f₁.op ≫ g₁.op)) r := by dsimp [a₀] - rw [NatTrans.naturality_apply, ha₁, Functor.map_comp, comp_apply] + rw [← RingCat.comp_apply, NatTrans.naturality, RingCat.comp_apply, ha₁, Functor.map_comp, + RingCat.comp_apply] have hb₀ : (φ.app (Opposite.op Z)) b₀ = (A.map (f₁.op ≫ g₁.op)) m := by dsimp [b₀] erw [NatTrans.naturality_apply, hb₁, Functor.map_comp, comp_apply] @@ -160,10 +163,9 @@ def SMulCandidate.mk' (S : Sieve X.unop) (hS : S ∈ J X.unop) dsimp at hg erw [← comp_apply, ← A.val.map_comp, ← NatTrans.naturality_apply, M₀.map_smul] refine (ha _ hg).trans (app_eq_of_isLocallyInjective α φ A.isSeparated _ _ _ _ ?_ ?_) - · rw [NatTrans.naturality_apply, ha₀] + · rw [← RingCat.comp_apply, NatTrans.naturality, RingCat.comp_apply, ha₀] apply (hr₀ _ hg).symm.trans - dsimp - rw [Functor.map_comp, comp_apply] + simp [RingCat.forget_map] · erw [NatTrans.naturality_apply, hb₀] apply (hm₀ _ hg).symm.trans dsimp @@ -253,7 +255,8 @@ protected lemma add_smul : smul α φ (r + r') m = smul α φ r m + smul α φ r refine J.intersection_covering (J.intersection_covering ?_ ?_) ?_ all_goals apply Presheaf.imageSieve_mem apply A.isSeparated _ _ hS - rintro Y f ⟨⟨⟨r₀ : R₀.obj _, hr₀⟩, ⟨r₀' : R₀.obj _, hr₀'⟩⟩, ⟨m₀, hm₀⟩⟩ + rintro Y f ⟨⟨⟨r₀ : R₀.obj _, (hr₀ : (α.app (Opposite.op Y)) r₀ = (R.val.map f.op) r)⟩, + ⟨r₀' : R₀.obj _, (hr₀' : (α.app (Opposite.op Y)) r₀' = (R.val.map f.op) r')⟩⟩, ⟨m₀, hm₀⟩⟩ rw [(A.val.map f.op).map_add, map_smul_eq α φ r m f.op r₀ hr₀ m₀ hm₀, map_smul_eq α φ r' m f.op r₀' hr₀' m₀ hm₀, map_smul_eq α φ (r + r') m f.op (r₀ + r₀') (by rw [map_add, map_add, hr₀, hr₀']) @@ -265,7 +268,9 @@ protected lemma mul_smul : smul α φ (r * r') m = smul α φ r (smul α φ r' m refine J.intersection_covering (J.intersection_covering ?_ ?_) ?_ all_goals apply Presheaf.imageSieve_mem apply A.isSeparated _ _ hS - rintro Y f ⟨⟨⟨r₀ : R₀.obj _, hr₀⟩, ⟨r₀' : R₀.obj _, hr₀'⟩⟩, ⟨m₀ : M₀.obj _, hm₀⟩⟩ + rintro Y f ⟨⟨⟨r₀ : R₀.obj _, (hr₀ : (α.app (Opposite.op Y)) r₀ = (R.val.map f.op) r)⟩, + ⟨r₀' : R₀.obj _, (hr₀' : (α.app (Opposite.op Y)) r₀' = (R.val.map f.op) r')⟩⟩, + ⟨m₀ : M₀.obj _, hm₀⟩⟩ rw [map_smul_eq α φ (r * r') m f.op (r₀ * r₀') (by rw [map_mul, map_mul, hr₀, hr₀']) m₀ hm₀, mul_smul, map_smul_eq α φ r (smul α φ r' m) f.op r₀ hr₀ (r₀' • m₀) @@ -291,9 +296,10 @@ lemma map_smul : apply J.intersection_covering all_goals apply Presheaf.imageSieve_mem apply A.isSeparated _ _ hS - rintro Y f ⟨⟨r₀, hr₀⟩, ⟨m₀, hm₀⟩⟩ + rintro Y f ⟨⟨r₀, + (hr₀ : (α.app (Opposite.op Y)).hom r₀ = (R.val.map f.op).hom ((R.val.map π).hom r))⟩, ⟨m₀, hm₀⟩⟩ rw [← comp_apply, ← Functor.map_comp, - map_smul_eq α φ r m (π ≫ f.op) r₀ (by rw [hr₀, Functor.map_comp, comp_apply]) m₀ + map_smul_eq α φ r m (π ≫ f.op) r₀ (by rw [hr₀, Functor.map_comp, RingCat.comp_apply]) m₀ (by rw [hm₀, Functor.map_comp, comp_apply]), map_smul_eq α φ (R.val.map π r) (A.val.map π m) f.op r₀ hr₀ m₀ hm₀] @@ -320,7 +326,7 @@ lemma toSheafify_app_apply (X : Cᵒᵖ) (x : M₀.obj X) : /-- `@[simp]`-normal form of `toSheafify_app_apply`. -/ @[simp] lemma toSheafify_app_apply' (X : Cᵒᵖ) (x : M₀.obj X) : - DFunLike.coe (F := (_ →ₗ[_] ↑((ModuleCat.restrictScalars (α.app X)).obj _))) + DFunLike.coe (F := (_ →ₗ[_] ↑((ModuleCat.restrictScalars (α.app X).hom).obj _))) ((toSheafify α φ).app X).hom x = φ.app X x := rfl @[simp] diff --git a/Mathlib/Algebra/Category/ModuleCat/Sheaf/ChangeOfRings.lean b/Mathlib/Algebra/Category/ModuleCat/Sheaf/ChangeOfRings.lean index a73bfc742c383..c715193fc39b7 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Sheaf/ChangeOfRings.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Sheaf/ChangeOfRings.lean @@ -53,7 +53,8 @@ noncomputable def restrictHomEquivOfIsLocallySurjective toFun f := (restrictScalars α).map f invFun g := homMk ((toPresheaf R).map g) (fun X r' m ↦ by apply hM₂.isSeparated _ _ (Presheaf.imageSieve_mem J α r') - rintro Y p ⟨r : R.obj _, hr⟩ + -- Type-ascript `hr` so it uses `RingCat.Hom.hom` instead of `ConcreteCategory.instFunLike` + rintro Y p ⟨r : R.obj _, (hr : α.app (Opposite.op Y) r = R'.map p.op r')⟩ have hg : ∀ (z : M₁.obj X), g.app _ (M₁.map p.op z) = M₂.map p.op (g.app X z) := fun z ↦ congr_fun ((forget _).congr_map (g.naturality p.op)) z change M₂.map p.op (g.app X (r' • m)) = M₂.map p.op (r' • show M₂.obj X from g.app X m) diff --git a/Mathlib/Algebra/Category/Ring/Adjunctions.lean b/Mathlib/Algebra/Category/Ring/Adjunctions.lean index bc0a55a174ab8..e4d4568695a11 100644 --- a/Mathlib/Algebra/Category/Ring/Adjunctions.lean +++ b/Mathlib/Algebra/Category/Ring/Adjunctions.lean @@ -27,11 +27,7 @@ polynomials with variables `x : X`. -/ def free : Type u ⥤ CommRingCat.{u} where obj α := of (MvPolynomial α ℤ) - map {X Y} f := (↑(rename f : _ →ₐ[ℤ] _) : MvPolynomial X ℤ →+* MvPolynomial Y ℤ) - -- TODO these next two fields can be done by `tidy`, but the calls in `dsimp` and `simp` it - -- generates are too slow. - map_id _ := RingHom.ext <| rename_id - map_comp f g := RingHom.ext fun p => (rename_rename f g p).symm + map {X Y} f := ofHom (↑(rename f : _ →ₐ[ℤ] _) : MvPolynomial X ℤ →+* MvPolynomial Y ℤ) @[simp] theorem free_obj_coe {α : Type u} : (free.obj α : Type u) = MvPolynomial α ℤ := @@ -44,13 +40,16 @@ theorem free_obj_coe {α : Type u} : (free.obj α : Type u) = MvPolynomial α theorem free_map_coe {α β : Type u} {f : α → β} : ⇑(free.map f) = ⇑(rename f) := rfl -/-- The free-forgetful adjunction for commutative rings. --/ +/-- The free-forgetful adjunction for commutative rings. -/ def adj : free ⊣ forget CommRingCat.{u} := Adjunction.mkOfHomEquiv - { homEquiv := fun _ _ => homEquiv + { homEquiv := fun _ _ ↦ + { toFun := fun f ↦ homEquiv f.hom + invFun := fun f ↦ ofHom <| homEquiv.symm f + left_inv := fun f ↦ congrArg ofHom (homEquiv.left_inv f.hom) + right_inv := fun f ↦ homEquiv.right_inv f } homEquiv_naturality_left_symm := fun {_ _ Y} f g => - RingHom.ext fun x => eval₂_cast_comp f (Int.castRingHom Y) g x } + hom_ext <| RingHom.ext fun x ↦ eval₂_cast_comp f (Int.castRingHom Y) g x } instance : (forget CommRingCat.{u}).IsRightAdjoint := ⟨_, ⟨adj⟩⟩ diff --git a/Mathlib/Algebra/Category/Ring/Basic.lean b/Mathlib/Algebra/Category/Ring/Basic.lean index 61b53a227b0c6..0933469109033 100644 --- a/Mathlib/Algebra/Category/Ring/Basic.lean +++ b/Mathlib/Algebra/Category/Ring/Basic.lean @@ -23,697 +23,629 @@ universe u v open CategoryTheory /-- The category of semirings. -/ -abbrev SemiRingCat : Type (u + 1) := - Bundled Semiring +structure SemiRingCat where + private mk :: + /-- The underlying type. -/ + carrier : Type u + [semiring : Semiring carrier] --- Porting note: typemax hack to fix universe complaints -/-- An alias for `Semiring.{max u v}`, to deal around unification issues. -/ -@[nolint checkUnivs] -abbrev SemiRingCatMax.{u1, u2} := SemiRingCat.{max u1 u2} +attribute [instance] SemiRingCat.semiring -namespace SemiRingCat +initialize_simps_projections SemiRingCat (-semiring) -/-- `RingHom` doesn't actually assume associativity. This alias is needed to make the category -theory machinery work. We use the same trick in `MonCat.AssocMonoidHom`. -/ -abbrev AssocRingHom (M N : Type*) [Semiring M] [Semiring N] := - RingHom M N +namespace SemiRingCat -instance bundledHom : BundledHom AssocRingHom where - toFun _ _ f := f - id _ := RingHom.id _ - comp _ _ _ f g := f.comp g +instance : CoeSort (SemiRingCat) (Type u) := + ⟨SemiRingCat.carrier⟩ --- Porting note: deriving fails for ConcreteCategory, adding instance manually. ---deriving instance LargeCategory, ConcreteCategory for SemiRingCat --- see https://github.com/leanprover-community/mathlib4/issues/5020 +attribute [coe] SemiRingCat.carrier -instance instSemiring (X : SemiRingCat) : Semiring X := X.str +/-- The object in the category of R-algebras associated to a type equipped with the appropriate +typeclasses. This is the preferred way to construct a term of `AlgebraCat R`. -/ +abbrev of (R : Type u) [Semiring R] : SemiRingCat := + ⟨R⟩ -instance instFunLike {X Y : SemiRingCat} : FunLike (X ⟶ Y) X Y := - ConcreteCategory.instFunLike +lemma coe_of (R : Type u) [Semiring R] : (of R : Type u) = R := + rfl --- Porting note (https://github.com/leanprover-community/mathlib4/pull/10754): added instance -instance instRingHomClass {X Y : SemiRingCat} : RingHomClass (X ⟶ Y) X Y := - RingHom.instRingHomClass +lemma of_carrier (R : SemiRingCat.{u}) : of R = R := rfl -lemma coe_id {X : SemiRingCat} : (𝟙 X : X → X) = id := rfl +variable {R} in +/-- The type of morphisms in `SemiRingCat`. -/ +@[ext] +structure Hom (R S : SemiRingCat) where + private mk :: + /-- The underlying ring hom. -/ + hom : R →+* S -lemma coe_comp {X Y Z : SemiRingCat} {f : X ⟶ Y} {g : Y ⟶ Z} : (f ≫ g : X → Z) = g ∘ f := rfl +instance : Category SemiRingCat where + Hom R S := Hom R S + id R := ⟨RingHom.id R⟩ + comp f g := ⟨g.hom.comp f.hom⟩ -@[simp] lemma forget_map {X Y : SemiRingCat} (f : X ⟶ Y) : - (forget SemiRingCat).map f = (f : X → Y) := rfl +instance {R S : SemiRingCat.{u}} : CoeFun (R ⟶ S) (fun _ ↦ R → S) where + coe f := f.hom -lemma ext {X Y : SemiRingCat} {f g : X ⟶ Y} (w : ∀ x : X, f x = g x) : f = g := - RingHom.ext w +@[simp] +lemma hom_id {R : SemiRingCat} : (𝟙 R : R ⟶ R).hom = RingHom.id R := rfl -/-- Construct a bundled SemiRing from the underlying type and typeclass. -/ -def of (R : Type u) [Semiring R] : SemiRingCat := - Bundled.of R +/- Provided for rewriting. -/ +lemma id_apply (R : SemiRingCat) (r : R) : + (𝟙 R : R ⟶ R) r = r := by simp @[simp] -theorem coe_of (R : Type u) [Semiring R] : (SemiRingCat.of R : Type u) = R := - rfl +lemma hom_comp {R S T : SemiRingCat} (f : R ⟶ S) (g : S ⟶ T) : + (f ≫ g).hom = g.hom.comp f.hom := rfl --- Coercing the identity morphism, as a ring homomorphism, gives the identity function. -@[simp] theorem coe_ringHom_id {X : SemiRingCat} : - @DFunLike.coe (X →+* X) X (fun _ ↦ X) _ (𝟙 X) = id := - rfl +/- Provided for rewriting. -/ +lemma comp_apply {R S T : SemiRingCat} (f : R ⟶ S) (g : S ⟶ T) (r : R) : + (f ≫ g) r = g (f r) := by simp --- Coercing `𝟙 (of X)` to a function should be expressed as the coercion of `RingHom.id X`. -@[simp] theorem coe_id_of {X : Type u} [Semiring X] : - @DFunLike.coe no_index (SemiRingCat.of X ⟶ SemiRingCat.of X) X - (fun _ ↦ X) _ - (𝟙 (of X)) = - @DFunLike.coe (X →+* X) X (fun _ ↦ X) _ (RingHom.id X) := - rfl +@[ext] +lemma hom_ext {R S : SemiRingCat} {f g : R ⟶ S} (hf : f.hom = g.hom) : f = g := + Hom.ext hf --- Coercing `f ≫ g`, where `f : of X ⟶ of Y` and `g : of Y ⟶ of Z`, to a function should be --- expressed in terms of the coercion of `g.comp f`. -@[simp] theorem coe_comp_of {X Y Z : Type u} [Semiring X] [Semiring Y] [Semiring Z] - (f : X →+* Y) (g : Y →+* Z) : - @DFunLike.coe no_index (SemiRingCat.of X ⟶ SemiRingCat.of Z) X - (fun _ ↦ Z) _ - (CategoryStruct.comp (X := SemiRingCat.of X) (Y := SemiRingCat.of Y) (Z := SemiRingCat.of Z) - f g) = - @DFunLike.coe (X →+* Z) X (fun _ ↦ Z) _ (RingHom.comp g f) := - rfl +/-- Typecheck a `RingHom` as a morphism in `SemiRingCat`. -/ +abbrev ofHom {R S : Type u} [Semiring R] [Semiring S] (f : R →+* S) : of R ⟶ of S := + ⟨f⟩ -/-- Variant of `SemiRingCat.coe_comp_of` for morphisms. -/ -@[simp] theorem coe_comp_of' {X Z : Type u} {Y : SemiRingCat.{u}} - [Semiring X] [Semiring Z] - (f : of X ⟶ Y) (g : Y ⟶ of Z) : - @DFunLike.coe no_index (of X ⟶ of Z) X (fun _ ↦ Z) _ (f ≫ g) = - @DFunLike.coe (X →+* Z) X (fun _ ↦ Z) _ (g.comp f) := by - rfl +lemma hom_ofHom {R S : Type u} [Semiring R] [Semiring S] (f : R →+* S) : (ofHom f).hom = f := rfl --- Sometimes neither the `ext` lemma for `SemiRingCat` nor for `RingHom` is applicable, --- because of incomplete unfolding of `SemiRingCat.of X ⟶ SemiRingCat.of Y := X →+* Y`, --- but this one will fire. -@[ext] theorem ext_of {X Y : Type u} [Semiring X] [Semiring Y] (f g : X →+* Y) - (h : ∀ x, f x = g x) : - @Eq (SemiRingCat.of X ⟶ SemiRingCat.of Y) f g := - RingHom.ext h +@[simp] +lemma ofHom_hom {R S : SemiRingCat} (f : R ⟶ S) : + ofHom (Hom.hom f) = f := rfl + +@[simp] +lemma ofHom_id {R : Type u} [Semiring R] : ofHom (RingHom.id R) = 𝟙 (of R) := rfl @[simp] -lemma RingEquiv_coe_eq {X Y : Type _} [Semiring X] [Semiring Y] (e : X ≃+* Y) : - (@DFunLike.coe (SemiRingCat.of X ⟶ SemiRingCat.of Y) _ (fun _ => (forget SemiRingCat).obj _) - ConcreteCategory.instFunLike (e : X →+* Y) : X → Y) = ↑e := +lemma ofHom_comp {R S T : Type u} [Semiring R] [Semiring S] [Semiring T] + (f : R →+* S) (g : S →+* T) : + ofHom (g.comp f) = ofHom f ≫ ofHom g := rfl +lemma ofHom_apply {R S : Type u} [Semiring R] [Semiring S] + (f : R →+* S) (r : R) : ofHom f r = f r := rfl + +@[simp] +lemma inv_hom_apply {R S : SemiRingCat} (e : R ≅ S) (r : R) : e.inv (e.hom r) = r := by + rw [← comp_apply] + simp + +@[simp] +lemma hom_inv_apply {R S : SemiRingCat} (e : R ≅ S) (s : S) : e.hom (e.inv s) = s := by + rw [← comp_apply] + simp + instance : Inhabited SemiRingCat := ⟨of PUnit⟩ -instance hasForgetToMonCat : HasForget₂ SemiRingCat MonCat := - BundledHom.mkHasForget₂ - (fun R hR => @MonoidWithZero.toMonoid R (@Semiring.toMonoidWithZero R hR)) - (fun {_ _} => RingHom.toMonoidHom) - (fun _ => rfl) +instance : ConcreteCategory.{u} SemiRingCat where + forget := + { obj := fun R => R + map := fun f => f.hom } + forget_faithful := ⟨fun h => by ext x; simpa using congrFun h x⟩ -instance hasForgetToAddCommMonCat : HasForget₂ SemiRingCat AddCommMonCat where - -- can't use BundledHom.mkHasForget₂, since AddCommMon is an induced category - forget₂ := - { obj := fun R => AddCommMonCat.of R - -- Porting note: This doesn't work without the `(_ := _)` trick. - map := fun {R₁ R₂} f => RingHom.toAddMonoidHom (α := R₁) (β := R₂) f } +lemma forget_obj {R : SemiRingCat} : (forget SemiRingCat).obj R = R := rfl -/-- Typecheck a `RingHom` as a morphism in `SemiRingCat`. -/ -def ofHom {R S : Type u} [Semiring R] [Semiring S] (f : R →+* S) : of R ⟶ of S := - f - --- Porting note: `simpNF` should not trigger on `rfl` lemmas. --- see https://github.com/leanprover/std4/issues/86 -@[simp, nolint simpNF] -theorem ofHom_apply {R S : Type u} [Semiring R] [Semiring S] (f : R →+* S) (x : R) : - ofHom f x = f x := +lemma forget_map {R S : SemiRingCat} (f : R ⟶ S) : + (forget SemiRingCat).map f = f := rfl -/-- A variant of `ofHom_apply` that makes `simpNF` happy -/ -@[simp] -theorem ofHom_apply' {R S : Type u} [Semiring R] [Semiring S] (f : R →+* S) (x : R) : - @DFunLike.coe no_index _ R (fun _ ↦ S) _ (ofHom f) x = f x := rfl +instance {R : SemiRingCat} : Semiring ((forget SemiRingCat).obj R) := + (inferInstance : Semiring R.carrier) -/-- -Ring equivalence are isomorphisms in category of semirings --/ +instance hasForgetToMonCat : HasForget₂ SemiRingCat MonCat where + forget₂ := + { obj := fun R ↦ MonCat.of R + map := fun f ↦ f.hom.toMonoidHom } + +instance hasForgetToAddCommMonCat : HasForget₂ SemiRingCat AddCommMonCat where + forget₂ := + { obj := fun R ↦ AddCommMonCat.of R + map := fun f ↦ f.hom.toAddMonoidHom } + +/-- Ring equivalence are isomorphisms in category of semirings -/ @[simps] -def _root_.RingEquiv.toSemiRingCatIso {X Y : Type u} [Semiring X] [Semiring Y] (e : X ≃+* Y) : - SemiRingCat.of X ≅ SemiRingCat.of Y where - hom := e.toRingHom - inv := e.symm.toRingHom +def _root_.RingEquiv.toSemiRingCatIso {R S : Type u} [Semiring R] [Semiring S] (e : R ≃+* S) : + of R ≅ of S where + hom := ⟨e⟩ + inv := ⟨e.symm⟩ instance forgetReflectIsos : (forget SemiRingCat).ReflectsIsomorphisms where reflects {X Y} f _ := by let i := asIso ((forget SemiRingCat).map f) - let ff : X →+* Y := f + let ff : X →+* Y := f.hom let e : X ≃+* Y := { ff, i.toEquiv with } exact e.toSemiRingCatIso.isIso_hom end SemiRingCat -/-- The category of rings. -/ -abbrev RingCat : Type (u + 1) := - Bundled Ring +/-- The category of semirings. -/ +structure RingCat where + private mk :: + /-- The underlying type. -/ + carrier : Type u + [ring : Ring carrier] -namespace RingCat +attribute [instance] RingCat.ring -instance : BundledHom.ParentProjection @Ring.toSemiring := - ⟨⟩ +initialize_simps_projections RingCat (-ring) --- Porting note: Another place where mathlib had derived a concrete category --- but this does not work here, so we add the instance manually. --- see https://github.com/leanprover-community/mathlib4/issues/5020 +namespace RingCat -instance (X : RingCat) : Ring X := X.str +instance : CoeSort (RingCat) (Type u) := + ⟨RingCat.carrier⟩ -instance instRing (X : RingCat) : Ring X := X.str +attribute [coe] RingCat.carrier -instance instFunLike {X Y : RingCat} : FunLike (X ⟶ Y) X Y := - -- Note: this is apparently _not_ defeq to RingHom.instFunLike with reducible transparency - ConcreteCategory.instFunLike +/-- The object in the category of R-algebras associated to a type equipped with the appropriate +typeclasses. This is the preferred way to construct a term of `AlgebraCat R`. -/ +abbrev of (R : Type u) [Ring R] : RingCat := + ⟨R⟩ --- Porting note (https://github.com/leanprover-community/mathlib4/pull/10754): added instance -instance instRingHomClass {X Y : RingCat} : RingHomClass (X ⟶ Y) X Y := - RingHom.instRingHomClass +lemma coe_of (R : Type u) [Ring R] : (of R : Type u) = R := + rfl -lemma coe_id {X : RingCat} : (𝟙 X : X → X) = id := rfl +lemma of_carrier (R : RingCat.{u}) : of R = R := rfl -lemma coe_comp {X Y Z : RingCat} {f : X ⟶ Y} {g : Y ⟶ Z} : (f ≫ g : X → Z) = g ∘ f := rfl +variable {R} in +/-- The type of morphisms in `RingCat`. -/ +@[ext] +structure Hom (R S : RingCat) where + private mk :: + /-- The underlying ring hom. -/ + hom : R →+* S -@[simp] lemma forget_map {X Y : RingCat} (f : X ⟶ Y) : (forget RingCat).map f = (f : X → Y) := rfl +instance : Category RingCat where + Hom R S := Hom R S + id R := ⟨RingHom.id R⟩ + comp f g := ⟨g.hom.comp f.hom⟩ -lemma ext {X Y : RingCat} {f g : X ⟶ Y} (w : ∀ x : X, f x = g x) : f = g := - RingHom.ext w +instance {R S : RingCat.{u}} : CoeFun (R ⟶ S) (fun _ ↦ R → S) where + coe f := f.hom -/-- Construct a bundled `RingCat` from the underlying type and typeclass. -/ -def of (R : Type u) [Ring R] : RingCat := - Bundled.of R +@[simp] +lemma hom_id {R : RingCat} : (𝟙 R : R ⟶ R).hom = RingHom.id R := rfl -/-- Typecheck a `RingHom` as a morphism in `RingCat`. -/ -def ofHom {R S : Type u} [Ring R] [Ring S] (f : R →+* S) : of R ⟶ of S := - f +/- Provided for rewriting. -/ +lemma id_apply (R : RingCat) (r : R) : + (𝟙 R : R ⟶ R) r = r := by simp -theorem ofHom_apply {R S : Type u} [Ring R] [Ring S] (f : R →+* S) (x : R) : ofHom f x = f x := - rfl +@[simp] +lemma hom_comp {R S T : RingCat} (f : R ⟶ S) (g : S ⟶ T) : + (f ≫ g).hom = g.hom.comp f.hom := rfl +/- Provided for rewriting. -/ +lemma comp_apply {R S T : RingCat} (f : R ⟶ S) (g : S ⟶ T) (r : R) : + (f ≫ g) r = g (f r) := by simp -instance : Inhabited RingCat := - ⟨of PUnit⟩ +@[ext] +lemma hom_ext {R S : RingCat} {f g : R ⟶ S} (hf : f.hom = g.hom) : f = g := + Hom.ext hf + +/-- Typecheck an `AlgHom` as a morphism in `AlgebraCat R`. -/ +abbrev ofHom {R S : Type u} [Ring R] [Ring S] (f : R →+* S) : of R ⟶ of S := + ⟨f⟩ -instance (R : RingCat) : Ring R := - R.str +lemma hom_ofHom {R S : Type u} [Ring R] [Ring S] (f : R →+* S) : (ofHom f).hom = f := rfl @[simp] -theorem coe_of (R : Type u) [Ring R] : (RingCat.of R : Type u) = R := - rfl +lemma ofHom_hom {R S : RingCat} (f : R ⟶ S) : + ofHom (Hom.hom f) = f := rfl -/-- A variant of `ofHom_apply` that makes `simpNF` happy -/ @[simp] -theorem ofHom_apply' {R S : Type u} [Ring R] [Ring S] (f : R →+* S) (x : R) : - @DFunLike.coe no_index _ R (fun _ ↦ S) _ (ofHom f) x = f x := rfl +lemma ofHom_id {R : Type u} [Ring R] : ofHom (RingHom.id R) = 𝟙 (of R) := rfl --- Coercing the identity morphism, as a ring homomorphism, gives the identity function. -@[simp] theorem coe_ringHom_id {X : RingCat} : - @DFunLike.coe (X →+* X) X (fun _ ↦ X) _ (𝟙 X) = id := +@[simp] +lemma ofHom_comp {R S T : Type u} [Ring R] [Ring S] [Ring T] + (f : R →+* S) (g : S →+* T) : + ofHom (g.comp f) = ofHom f ≫ ofHom g := rfl --- Coercing `𝟙 (of X)` to a function should be expressed as the coercion of `RingHom.id X`. -@[simp] theorem coe_id_of {X : Type u} [Ring X] : - @DFunLike.coe no_index (RingCat.of X ⟶ RingCat.of X) X - (fun _ ↦ X) _ - (𝟙 (of X)) = - @DFunLike.coe (X →+* X) X (fun _ ↦ X) _ (RingHom.id X) := - rfl +lemma ofHom_apply {R S : Type u} [Ring R] [Ring S] + (f : R →+* S) (r : R) : ofHom f r = f r := rfl --- Coercing `f ≫ g`, where `f : of X ⟶ of Y` and `g : of Y ⟶ of Z`, to a function should be --- expressed in terms of the coercion of `g.comp f`. -@[simp] theorem coe_comp_of {X Y Z : Type u} [Ring X] [Ring Y] [Ring Z] - (f : X →+* Y) (g : Y →+* Z) : - @DFunLike.coe no_index (RingCat.of X ⟶ RingCat.of Z) X - (fun _ ↦ Z) _ - (CategoryStruct.comp (X := RingCat.of X) (Y := RingCat.of Y) (Z := RingCat.of Z) - f g) = - @DFunLike.coe (X →+* Z) X (fun _ ↦ Z) _ (RingHom.comp g f) := - rfl +@[simp] +lemma inv_hom_apply {R S : RingCat} (e : R ≅ S) (r : R) : e.inv (e.hom r) = r := by + rw [← comp_apply] + simp -/-- Variant of `RingCat.coe_comp_of` for morphisms. -/ -@[simp] theorem coe_comp_of' {X Z : Type u} {Y : RingCat.{u}} [Ring X] [Ring Z] - (f : of X ⟶ Y) (g : Y ⟶ of Z) : - @DFunLike.coe no_index (of X ⟶ of Z) X (fun _ ↦ Z) _ (f ≫ g) = - @DFunLike.coe (X →+* Z) X (fun _ ↦ Z) _ (g.comp f) := by - rfl +@[simp] +lemma hom_inv_apply {R S : RingCat} (e : R ≅ S) (s : S) : e.hom (e.inv s) = s := by + rw [← comp_apply] + simp --- Sometimes neither the `ext` lemma for `RingCat` nor for `RingHom` is applicable, --- because of incomplete unfolding of `RingCat.of X ⟶ RingCat.of Y := X →+* Y`, --- but this one will fire. -@[ext] theorem ext_of {X Y : Type u} [Ring X] [Ring Y] (f g : X →+* Y) - (h : ∀ x, f x = g x) : - @Eq (RingCat.of X ⟶ RingCat.of Y) f g := - RingHom.ext h +instance : Inhabited RingCat := + ⟨of PUnit⟩ -@[simp] -lemma RingEquiv_coe_eq {X Y : Type _} [Ring X] [Ring Y] (e : X ≃+* Y) : - (@DFunLike.coe (RingCat.of X ⟶ RingCat.of Y) _ (fun _ => (forget RingCat).obj _) - ConcreteCategory.instFunLike (e : X →+* Y) : X → Y) = ↑e := +instance : ConcreteCategory.{u} RingCat where + forget := + { obj := fun R => R + map := fun f => f.hom } + forget_faithful := ⟨fun h => by ext x; simpa using congrFun h x⟩ + +lemma forget_obj {R : RingCat} : (forget RingCat).obj R = R := rfl + +lemma forget_map {R S : RingCat} (f : R ⟶ S) : + (forget RingCat).map f = f := rfl -instance hasForgetToSemiRingCat : HasForget₂ RingCat SemiRingCat := - BundledHom.forget₂ _ _ +instance {R : RingCat} : Ring ((forget RingCat).obj R) := + (inferInstance : Ring R.carrier) + +instance hasForgetToSemiRingCat : HasForget₂ RingCat SemiRingCat where + forget₂ := + { obj := fun R ↦ SemiRingCat.of R + map := fun f ↦ SemiRingCat.ofHom f.hom } instance hasForgetToAddCommGrp : HasForget₂ RingCat AddCommGrp where - -- can't use BundledHom.mkHasForget₂, since AddCommGroup is an induced category forget₂ := - { obj := fun R => AddCommGrp.of R - -- Porting note: use `(_ := _)` similar to above. - map := fun {R₁ R₂} f => RingHom.toAddMonoidHom (α := R₁) (β := R₂) f } + { obj := fun R ↦ AddCommGrp.of R + map := fun f ↦ f.hom.toAddMonoidHom } + +/-- Ring equivalence are isomorphisms in category of semirings -/ +@[simps] +def _root_.RingEquiv.toRingCatIso {R S : Type u} [Ring R] [Ring S] (e : R ≃+* S) : + of R ≅ of S where + hom := ⟨e⟩ + inv := ⟨e.symm⟩ + +instance forgetReflectIsos : (forget RingCat).ReflectsIsomorphisms where + reflects {X Y} f _ := by + let i := asIso ((forget RingCat).map f) + let ff : X →+* Y := f.hom + let e : X ≃+* Y := { ff, i.toEquiv with } + exact e.toRingCatIso.isIso_hom end RingCat -/-- The category of commutative semirings. -/ -abbrev CommSemiRingCat : Type (u + 1) := - Bundled CommSemiring +/-- The category of semirings. -/ +structure CommSemiRingCat where + private mk :: + /-- The underlying type. -/ + carrier : Type u + [commSemiring : CommSemiring carrier] + +attribute [instance] CommSemiRingCat.commSemiring + +initialize_simps_projections CommSemiRingCat (-commSemiring) namespace CommSemiRingCat -instance : BundledHom.ParentProjection @CommSemiring.toSemiring := - ⟨⟩ +instance : CoeSort (CommSemiRingCat) (Type u) := + ⟨CommSemiRingCat.carrier⟩ --- Porting note: again, deriving fails for concrete category instances. --- see https://github.com/leanprover-community/mathlib4/issues/5020 -deriving instance LargeCategory for CommSemiRingCat +attribute [coe] CommSemiRingCat.carrier -instance : ConcreteCategory CommSemiRingCat := by - dsimp [CommSemiRingCat] - infer_instance +/-- The object in the category of R-algebras associated to a type equipped with the appropriate +typeclasses. This is the preferred way to construct a term of `AlgebraCat R`. -/ +abbrev of (R : Type u) [CommSemiring R] : CommSemiRingCat := + ⟨R⟩ -instance : CoeSort CommSemiRingCat Type* where - coe X := X.α +lemma coe_of (R : Type u) [CommSemiring R] : (of R : Type u) = R := + rfl -instance (X : CommSemiRingCat) : CommSemiring X := X.str +lemma of_carrier (R : CommSemiRingCat.{u}) : of R = R := rfl -instance instCommSemiring (X : CommSemiRingCat) : CommSemiring X := X.str +variable {R} in +/-- The type of morphisms in `CommSemiRingCat`. -/ +@[ext] +structure Hom (R S : CommSemiRingCat) where + private mk :: + /-- The underlying ring hom. -/ + hom : R →+* S -instance instCommSemiring' (X : CommSemiRingCat) : CommSemiring <| (forget CommSemiRingCat).obj X := - X.str +instance : Category CommSemiRingCat where + Hom R S := Hom R S + id R := ⟨RingHom.id R⟩ + comp f g := ⟨g.hom.comp f.hom⟩ -instance instFunLike {X Y : CommSemiRingCat} : FunLike (X ⟶ Y) X Y := - -- Note: this is apparently _not_ defeq to RingHom.instFunLike with reducible transparency - ConcreteCategory.instFunLike +instance {R S : CommSemiRingCat.{u}} : CoeFun (R ⟶ S) (fun _ ↦ R → S) where + coe f := f.hom --- Porting note (https://github.com/leanprover-community/mathlib4/pull/10754): added instance -instance instRingHomClass {X Y : CommSemiRingCat} : RingHomClass (X ⟶ Y) X Y := - RingHom.instRingHomClass +@[simp] +lemma hom_id {R : CommSemiRingCat} : (𝟙 R : R ⟶ R).hom = RingHom.id R := rfl -lemma coe_id {X : CommSemiRingCat} : (𝟙 X : X → X) = id := rfl +/- Provided for rewriting. -/ +lemma id_apply (R : CommSemiRingCat) (r : R) : + (𝟙 R : R ⟶ R) r = r := by simp -lemma coe_comp {X Y Z : CommSemiRingCat} {f : X ⟶ Y} {g : Y ⟶ Z} : (f ≫ g : X → Z) = g ∘ f := rfl +@[simp] +lemma hom_comp {R S T : CommSemiRingCat} (f : R ⟶ S) (g : S ⟶ T) : + (f ≫ g).hom = g.hom.comp f.hom := rfl -@[simp] lemma forget_map {X Y : CommSemiRingCat} (f : X ⟶ Y) : - (forget CommSemiRingCat).map f = (f : X → Y) := rfl +/- Provided for rewriting. -/ +lemma comp_apply {R S T : CommSemiRingCat} (f : R ⟶ S) (g : S ⟶ T) (r : R) : + (f ≫ g) r = g (f r) := by simp -lemma ext {X Y : CommSemiRingCat} {f g : X ⟶ Y} (w : ∀ x : X, f x = g x) : f = g := - RingHom.ext w +@[ext] +lemma hom_ext {R S : CommSemiRingCat} {f g : R ⟶ S} (hf : f.hom = g.hom) : f = g := + Hom.ext hf -/-- Construct a bundled `CommSemiRingCat` from the underlying type and typeclass. -/ -def of (R : Type u) [CommSemiring R] : CommSemiRingCat := - Bundled.of R +/-- Typecheck an `AlgHom` as a morphism in `AlgebraCat R`. -/ +abbrev ofHom {R S : Type u} [CommSemiring R] [CommSemiring S] (f : R →+* S) : of R ⟶ of S := + ⟨f⟩ -/-- Typecheck a `RingHom` as a morphism in `CommSemiRingCat`. -/ -def ofHom {R S : Type u} [CommSemiring R] [CommSemiring S] (f : R →+* S) : of R ⟶ of S := - f +lemma hom_ofHom {R S : Type u} [CommSemiring R] [CommSemiring S] (f : R →+* S) : + (ofHom f).hom = f := rfl @[simp] -lemma RingEquiv_coe_eq {X Y : Type _} [CommSemiring X] [CommSemiring Y] (e : X ≃+* Y) : - (@DFunLike.coe (CommSemiRingCat.of X ⟶ CommSemiRingCat.of Y) _ - (fun _ => (forget CommSemiRingCat).obj _) - ConcreteCategory.instFunLike (e : X →+* Y) : X → Y) = ↑e := - rfl +lemma ofHom_hom {R S : CommSemiRingCat} (f : R ⟶ S) : + ofHom (Hom.hom f) = f := rfl -theorem ofHom_apply {R S : Type u} [CommSemiring R] [CommSemiring S] (f : R →+* S) (x : R) : - ofHom f x = f x := - rfl +@[simp] +lemma ofHom_id {R : Type u} [CommSemiring R] : ofHom (RingHom.id R) = 𝟙 (of R) := rfl -/-- A variant of `ofHom_apply` that makes `simpNF` happy -/ @[simp] -theorem ofHom_apply' {R S : Type u} [CommSemiring R] [CommSemiring S] (f : R →+* S) (x : R) : - @DFunLike.coe no_index _ R (fun _ ↦ S) _ (ofHom f) x = f x := rfl +lemma ofHom_comp {R S T : Type u} [CommSemiring R] [CommSemiring S] [CommSemiring T] + (f : R →+* S) (g : S →+* T) : + ofHom (g.comp f) = ofHom f ≫ ofHom g := + rfl -instance : Inhabited CommSemiRingCat := - ⟨of PUnit⟩ +lemma ofHom_apply {R S : Type u} [CommSemiring R] [CommSemiring S] + (f : R →+* S) (r : R) : ofHom f r = f r := rfl -instance (R : CommSemiRingCat) : CommSemiring R := - R.str +@[simp] +lemma inv_hom_apply {R S : CommSemiRingCat} (e : R ≅ S) (r : R) : e.inv (e.hom r) = r := by + rw [← comp_apply] + simp @[simp] -theorem coe_of (R : Type u) [CommSemiring R] : (CommSemiRingCat.of R : Type u) = R := - rfl +lemma hom_inv_apply {R S : CommSemiRingCat} (e : R ≅ S) (s : S) : e.hom (e.inv s) = s := by + rw [← comp_apply] + simp --- Coercing the identity morphism, as a ring homomorphism, gives the identity function. -@[simp] theorem coe_ringHom_id {X : CommSemiRingCat} : - @DFunLike.coe (X →+* X) X (fun _ ↦ X) _ (𝟙 X) = id := - rfl +instance : Inhabited CommSemiRingCat := + ⟨of PUnit⟩ --- Coercing `𝟙 (of X)` to a function should be expressed as the coercion of `RingHom.id X`. -@[simp] theorem coe_id_of {X : Type u} [CommSemiring X] : - @DFunLike.coe no_index (CommSemiRingCat.of X ⟶ CommSemiRingCat.of X) X - (fun _ ↦ X) _ - (𝟙 (of X)) = - @DFunLike.coe (X →+* X) X (fun _ ↦ X) _ (RingHom.id X) := - rfl +instance : ConcreteCategory.{u} CommSemiRingCat where + forget := + { obj := fun R => R + map := fun f => f.hom } + forget_faithful := ⟨fun h => by ext x; simpa using congrFun h x⟩ --- Coercing `f ≫ g`, where `f : of X ⟶ of Y` and `g : of Y ⟶ of Z`, to a function should be --- expressed in terms of the coercion of `g.comp f`. -@[simp] theorem coe_comp_of {X Y Z : Type u} [CommSemiring X] [CommSemiring Y] [CommSemiring Z] - (f : X →+* Y) (g : Y →+* Z) : - @DFunLike.coe no_index (CommSemiRingCat.of X ⟶ CommSemiRingCat.of Z) X - (fun _ ↦ Z) _ - (CategoryStruct.comp (X := CommSemiRingCat.of X) (Y := CommSemiRingCat.of Y) - (Z := CommSemiRingCat.of Z) f g) = - @DFunLike.coe (X →+* Z) X (fun _ ↦ Z) _ (RingHom.comp g f) := - rfl +lemma forget_obj {R : CommSemiRingCat} : (forget CommSemiRingCat).obj R = R := rfl -/-- Variant of `CommSemiRingCat.coe_comp_of` for morphisms. -/ -@[simp] theorem coe_comp_of' {X Z : Type u} {Y : CommSemiRingCat.{u}} - [CommSemiring X] [CommSemiring Z] - (f : of X ⟶ Y) (g : Y ⟶ of Z) : - @DFunLike.coe no_index (of X ⟶ of Z) X (fun _ ↦ Z) _ (f ≫ g) = - @DFunLike.coe (X →+* Z) X (fun _ ↦ Z) _ (g.comp f) := by +lemma forget_map {R S : CommSemiRingCat} (f : R ⟶ S) : + (forget CommSemiRingCat).map f = f := rfl --- Sometimes neither the `ext` lemma for `CommSemiRingCat` nor for `RingHom` is applicable, --- because of incomplete unfolding of `CommSemiRingCat.of X ⟶ CommSemiRingCat.of Y := X →+* Y`, --- but this one will fire. -@[ext] theorem ext_of {X Y : Type u} [CommSemiring X] [CommSemiring Y] (f g : X →+* Y) - (h : ∀ x, f x = g x) : - @Eq (CommSemiRingCat.of X ⟶ CommSemiRingCat.of Y) f g := - RingHom.ext h +instance {R : CommSemiRingCat} : CommSemiring ((forget CommSemiRingCat).obj R) := + (inferInstance : CommSemiring R.carrier) -instance hasForgetToSemiRingCat : HasForget₂ CommSemiRingCat SemiRingCat := - BundledHom.forget₂ _ _ +instance hasForgetToSemiRingCat : HasForget₂ CommSemiRingCat SemiRingCat where + forget₂ := + { obj := fun R ↦ ⟨R⟩ + map := fun f ↦ ⟨f.hom⟩ } /-- The forgetful functor from commutative rings to (multiplicative) commutative monoids. -/ -instance hasForgetToCommMonCat : HasForget₂ CommSemiRingCat CommMonCat := - HasForget₂.mk' (fun R : CommSemiRingCat => CommMonCat.of R) (fun _ => rfl) - -- Porting note: `(_ := _)` trick - (fun {R₁ R₂} f => RingHom.toMonoidHom (α := R₁) (β := R₂) f) (by rfl) +instance hasForgetToCommMonCat : HasForget₂ CommSemiRingCat CommMonCat where + forget₂ := + { obj := fun R ↦ CommMonCat.of R + map := fun f ↦ f.hom.toMonoidHom } -/-- -Ring equivalence are isomorphisms in category of commutative semirings --/ +/-- Ring equivalence are isomorphisms in category of semirings -/ @[simps] -def _root_.RingEquiv.toCommSemiRingCatIso {X Y : Type u} [CommSemiring X] [CommSemiring Y] - (e : X ≃+* Y) : CommSemiRingCat.of X ≅ CommSemiRingCat.of Y where - hom := e.toRingHom - inv := e.symm.toRingHom +def _root_.RingEquiv.toCommSemiRingCatIso + {R S : Type u} [CommSemiring R] [CommSemiring S] (e : R ≃+* S) : + of R ≅ of S where + hom := ⟨e⟩ + inv := ⟨e.symm⟩ instance forgetReflectIsos : (forget CommSemiRingCat).ReflectsIsomorphisms where reflects {X Y} f _ := by let i := asIso ((forget CommSemiRingCat).map f) - let ff : X →+* Y := f + let ff : X →+* Y := f.hom let e : X ≃+* Y := { ff, i.toEquiv with } - exact ⟨e.toSemiRingCatIso.isIso_hom.1⟩ + exact e.toCommSemiRingCatIso.isIso_hom end CommSemiRingCat -/-- The category of commutative rings. -/ -abbrev CommRingCat : Type (u + 1) := - Bundled CommRing +/-- The category of semirings. -/ +structure CommRingCat where + private mk :: + /-- The underlying type. -/ + carrier : Type u + [commRing : CommRing carrier] -namespace CommRingCat +attribute [instance] CommRingCat.commRing -instance : BundledHom.ParentProjection @CommRing.toRing := - ⟨⟩ +initialize_simps_projections CommRingCat (-commRing) --- Porting note: deriving fails for concrete category. --- see https://github.com/leanprover-community/mathlib4/issues/5020 -deriving instance LargeCategory for CommRingCat +namespace CommRingCat -instance instCommRing (X : CommRingCat) : CommRing X := X.str +instance : CoeSort (CommRingCat) (Type u) := + ⟨CommRingCat.carrier⟩ -instance instCommRing' (X : CommRingCat) : CommRing <| (forget CommRingCat).obj X := X.str +attribute [coe] CommRingCat.carrier -instance instFunLike {X Y : CommRingCat} : FunLike (X ⟶ Y) X Y := - -- Note: this is apparently _not_ defeq to RingHom.instFunLike with reducible transparency - ConcreteCategory.instFunLike +/-- The object in the category of R-algebras associated to a type equipped with the appropriate +typeclasses. This is the preferred way to construct a term of `AlgebraCat R`. -/ +abbrev of (R : Type u) [CommRing R] : CommRingCat := + ⟨R⟩ --- Porting note (https://github.com/leanprover-community/mathlib4/pull/10754): added instance -instance instRingHomClass {X Y : CommRingCat} : RingHomClass (X ⟶ Y) X Y := - RingHom.instRingHomClass +lemma coe_of (R : Type u) [CommRing R] : (of R : Type u) = R := + rfl -lemma coe_id {X : CommRingCat} : (𝟙 X : X → X) = id := rfl +lemma of_carrier (R : CommRingCat.{u}) : of R = R := rfl -lemma coe_comp {X Y Z : CommRingCat} {f : X ⟶ Y} {g : Y ⟶ Z} : (f ≫ g : X → Z) = g ∘ f := rfl +variable {R} in +/-- The type of morphisms in `CommRingCat`. -/ +@[ext] +structure Hom (R S : CommRingCat) where + private mk :: + /-- The underlying ring hom. -/ + hom : R →+* S -/-- Specialization of `ConcreteCategory.id_apply` because `simp` can't see through the defeq. -/ -@[simp] lemma id_apply (R : CommRingCat) (x : R) : 𝟙 R x = x := rfl +instance : Category CommRingCat where + Hom R S := Hom R S + id R := ⟨RingHom.id R⟩ + comp f g := ⟨g.hom.comp f.hom⟩ + +instance {R S : CommRingCat.{u}} : CoeFun (R ⟶ S) (fun _ ↦ R → S) where + coe f := f.hom @[simp] -theorem comp_apply {R S T : CommRingCat} (f : R ⟶ S) (g : S ⟶ T) (x : R) : - (f ≫ g) x = g (f x) := rfl +lemma hom_id {R : CommRingCat} : (𝟙 R : R ⟶ R).hom = RingHom.id R := rfl -@[simp] theorem forget_obj (R : CommRingCat) : (forget _).obj R = R := rfl +/- Provided for rewriting. -/ +lemma id_apply (R : CommRingCat) (r : R) : + (𝟙 R : R ⟶ R) r = r := by simp -@[simp] lemma forget_map {X Y : CommRingCat} (f : X ⟶ Y) : - (forget CommRingCat).map f = (f : X → Y) := rfl +@[simp] +lemma hom_comp {R S T : CommRingCat} (f : R ⟶ S) (g : S ⟶ T) : + (f ≫ g).hom = g.hom.comp f.hom := rfl -lemma ext {X Y : CommRingCat} {f g : X ⟶ Y} (w : ∀ x : X, f x = g x) : f = g := - RingHom.ext w +/- Provided for rewriting. -/ +lemma comp_apply {R S T : CommRingCat} (f : R ⟶ S) (g : S ⟶ T) (r : R) : + (f ≫ g) r = g (f r) := by simp -/-- Construct a bundled `CommRingCat` from the underlying type and typeclass. -/ -def of (R : Type u) [CommRing R] : CommRingCat := - Bundled.of R +@[ext] +lemma hom_ext {R S : CommRingCat} {f g : R ⟶ S} (hf : f.hom = g.hom) : f = g := + Hom.ext hf -instance instFunLike' {X : Type*} [CommRing X] {Y : CommRingCat} : - FunLike (CommRingCat.of X ⟶ Y) X Y := - -- Note: this is apparently _not_ defeq to RingHom.instFunLike with reducible transparency - ConcreteCategory.instFunLike +/-- Typecheck an `AlgHom` as a morphism in `AlgebraCat R`. -/ +abbrev ofHom {R S : Type u} [CommRing R] [CommRing S] (f : R →+* S) : of R ⟶ of S := + ⟨f⟩ -instance instFunLike'' {X : CommRingCat} {Y : Type*} [CommRing Y] : - FunLike (X ⟶ CommRingCat.of Y) X Y := - -- Note: this is apparently _not_ defeq to RingHom.instFunLike with reducible transparency - ConcreteCategory.instFunLike +lemma hom_ofHom {R S : Type u} [CommRing R] [CommRing S] (f : R →+* S) : + (ofHom f).hom = f := rfl -instance instFunLike''' {X Y : Type _} [CommRing X] [CommRing Y] : - FunLike (CommRingCat.of X ⟶ CommRingCat.of Y) X Y := - -- Note: this is apparently _not_ defeq to RingHom.instFunLike with reducible transparency - ConcreteCategory.instFunLike +@[simp] +lemma ofHom_hom {R S : CommRingCat} (f : R ⟶ S) : + ofHom (Hom.hom f) = f := rfl -/-- Typecheck a `RingHom` as a morphism in `CommRingCat`. -/ -def ofHom {R S : Type u} [CommRing R] [CommRing S] (f : R →+* S) : of R ⟶ of S := - f +@[simp] +lemma ofHom_id {R : Type u} [CommRing R] : ofHom (RingHom.id R) = 𝟙 (of R) := rfl @[simp] -lemma RingEquiv_coe_eq {X Y : Type _} [CommRing X] [CommRing Y] (e : X ≃+* Y) : - (@DFunLike.coe (CommRingCat.of X ⟶ CommRingCat.of Y) X (fun _ => Y) - ConcreteCategory.instFunLike (e : X →+* Y) : X → Y) = ↑e := +lemma ofHom_comp {R S T : Type u} [CommRing R] [CommRing S] [CommRing T] + (f : R →+* S) (g : S →+* T) : + ofHom (g.comp f) = ofHom f ≫ ofHom g := rfl -theorem ofHom_apply {R S : Type u} [CommRing R] [CommRing S] (f : R →+* S) (x : R) : - ofHom f x = f x := - rfl +lemma ofHom_apply {R S : Type u} [CommRing R] [CommRing S] + (f : R →+* S) (r : R) : ofHom f r = f r := rfl + +@[simp] +lemma inv_hom_apply {R S : CommRingCat} (e : R ≅ S) (r : R) : e.inv (e.hom r) = r := by + rw [← comp_apply] + simp -/-- A variant of `ofHom_apply` that makes `simpNF` happy -/ @[simp] -theorem ofHom_apply' {R S : Type u} [CommRing R] [CommRing S] (f : R →+* S) (x : R) : - @DFunLike.coe no_index _ R (fun _ ↦ S) _ (ofHom f) x = f x := rfl +lemma hom_inv_apply {R S : CommRingCat} (e : R ≅ S) (s : S) : e.hom (e.inv s) = s := by + rw [← comp_apply] + simp instance : Inhabited CommRingCat := ⟨of PUnit⟩ -instance (R : CommRingCat) : CommRing R := - R.str +instance : ConcreteCategory.{u} CommRingCat where + forget := + { obj := fun R => R + map := fun f => f.hom } + forget_faithful := ⟨fun h => by ext x; simpa using congrFun h x⟩ -@[simp] -theorem coe_of (R : Type u) [CommRing R] : (CommRingCat.of R : Type u) = R := - rfl +lemma forget_obj {R : CommRingCat} : (forget CommRingCat).obj R = R := rfl --- Coercing the identity morphism, as a ring homomorphism, gives the identity function. -@[simp] theorem coe_ringHom_id {X : CommRingCat} : - @DFunLike.coe (X →+* X) X (fun _ ↦ X) _ (𝟙 X) = id := +lemma forget_map {R S : CommRingCat} (f : R ⟶ S) : + (forget CommRingCat).map f = f := rfl --- Coercing `𝟙 (of X)` to a function should be expressed as the coercion of `RingHom.id X`. -@[simp] theorem coe_id_of {X : Type u} [CommRing X] : - @DFunLike.coe no_index (CommRingCat.of X ⟶ CommRingCat.of X) X - (fun _ ↦ X) _ - (𝟙 (of X)) = - @DFunLike.coe (X →+* X) X (fun _ ↦ X) _ (RingHom.id X) := - rfl +instance {R : CommRingCat} : CommRing ((forget CommRingCat).obj R) := + (inferInstance : CommRing R.carrier) --- Coercing `f ≫ g`, where `f : of X ⟶ of Y` and `g : of Y ⟶ of Z`, to a function should be --- expressed in terms of the coercion of `g.comp f`. -@[simp] theorem coe_comp_of {X Y Z : Type u} [CommRing X] [CommRing Y] [CommRing Z] - (f : X →+* Y) (g : Y →+* Z) : - @DFunLike.coe no_index (CommRingCat.of X ⟶ CommRingCat.of Z) X - (fun _ ↦ Z) _ - (CategoryStruct.comp (X := CommRingCat.of X) (Y := CommRingCat.of Y) (Z := CommRingCat.of Z) - f g) = - @DFunLike.coe (X →+* Z) X (fun _ ↦ Z) _ (RingHom.comp g f) := - rfl +instance hasForgetToRingCat : HasForget₂ CommRingCat RingCat where + forget₂ := + { obj := fun R ↦ RingCat.of R + map := fun f ↦ RingCat.ofHom f.hom } -/-- Variant of `CommRingCat.coe_comp_of` for morphisms. -/ -@[simp] theorem coe_comp_of' {X Z : Type u} {Y : CommRingCat.{u}} [CommRing X] [CommRing Z] - (f : of X ⟶ Y) (g : Y ⟶ of Z) : - @DFunLike.coe no_index (of X ⟶ of Z) X (fun _ ↦ Z) _ (f ≫ g) = - @DFunLike.coe (X →+* Z) X (fun _ ↦ Z) _ (g.comp f) := by +@[simp] lemma forgetToRingCat_map_hom {R S : CommRingCat} (f : R ⟶ S) : + ((forget₂ CommRingCat RingCat).map f).hom = f.hom := rfl --- Sometimes neither the `ext` lemma for `CommRingCat` nor for `RingHom` is applicable, --- because of incomplete unfolding of `CommRingCat.of X ⟶ CommRingCat.of Y := X →+* Y`, --- but this one will fire. -@[ext] theorem ext_of {X Y : Type u} [CommRing X] [CommRing Y] (f g : X →+* Y) - (h : ∀ x, f x = g x) : - @Eq (CommRingCat.of X ⟶ CommRingCat.of Y) f g := - RingHom.ext h - -instance hasForgetToRingCat : HasForget₂ CommRingCat RingCat := - BundledHom.forget₂ _ _ - -@[simp] lemma forgetToRingCat_obj (A : CommRingCat.{u}) : - ((forget₂ _ RingCat).obj A : Type _) = A := rfl +@[simp] lemma forgetToRingCat_obj {R : CommRingCat} : + (((forget₂ CommRingCat RingCat).obj R) : Type u) = R := + rfl -@[simp] lemma forgetToRingCat_map_apply {A B : CommRingCat.{u}} (f : A ⟶ B) (a : A) : - DFunLike.coe (α := A) (β := fun _ ↦ B) ((forget₂ _ RingCat).map f) a = f a := rfl +instance hasForgetToAddCommMonCat : HasForget₂ CommRingCat CommSemiRingCat where + forget₂ := + { obj := fun R ↦ CommSemiRingCat.of R + map := fun f ↦ CommSemiRingCat.ofHom f.hom } -/-- The forgetful functor from commutative rings to (multiplicative) commutative monoids. -/ -instance hasForgetToCommSemiRingCat : HasForget₂ CommRingCat CommSemiRingCat := - HasForget₂.mk' (fun R : CommRingCat => CommSemiRingCat.of R) (fun _ => rfl) - (fun {_ _} f => f) (by rfl) +/-- Ring equivalence are isomorphisms in category of semirings -/ +@[simps] +def _root_.RingEquiv.toCommRingCatIso + {R S : Type u} [CommRing R] [CommRing S] (e : R ≃+* S) : + of R ≅ of S where + hom := ⟨e⟩ + inv := ⟨e.symm⟩ -instance : (forget₂ CommRingCat CommSemiRingCat).Full where map_surjective f := ⟨f, rfl⟩ +instance forgetReflectIsos : (forget CommRingCat).ReflectsIsomorphisms where + reflects {X Y} f _ := by + let i := asIso ((forget CommRingCat).map f) + let ff : X →+* Y := f.hom + let e : X ≃+* Y := { ff, i.toEquiv with } + exact e.toCommRingCatIso.isIso_hom end CommRingCat --- We verify that simp lemmas apply when coercing morphisms to functions. -example {R S : CommRingCat} (i : R ⟶ S) (r : R) (h : r = 0) : i r = 0 := by simp [h] - -namespace RingEquiv +namespace CategoryTheory.Iso -variable {X Y : Type u} +/-- Build a `RingEquiv` from an isomorphism in the category `SemiRingCat`. -/ +def semiRingCatIsoToRingEquiv {R S : SemiRingCat.{u}} (e : R ≅ S) : R ≃+* S := + RingEquiv.ofHomInv e.hom.hom e.inv.hom (by ext; simp) (by ext; simp) -/-- Build an isomorphism in the category `RingCat` from a `RingEquiv` between `RingCat`s. -/ -@[simps] -def toRingCatIso [Ring X] [Ring Y] (e : X ≃+* Y) : RingCat.of X ≅ RingCat.of Y where - hom := e.toRingHom - inv := e.symm.toRingHom +/-- Build a `RingEquiv` from an isomorphism in the category `RingCat`. -/ +def ringCatIsoToRingEquiv {R S : RingCat.{u}} (e : R ≅ S) : R ≃+* S := + RingEquiv.ofHomInv e.hom.hom e.inv.hom (by ext; simp) (by ext; simp) -/-- Build an isomorphism in the category `CommRingCat` from a `RingEquiv` between `CommRingCat`s. -/ -@[simps] -def toCommRingCatIso [CommRing X] [CommRing Y] (e : X ≃+* Y) : - CommRingCat.of X ≅ CommRingCat.of Y where - hom := e.toRingHom - inv := e.symm.toRingHom +/-- Build a `RingEquiv` from an isomorphism in the category `CommSemiRingCat`. -/ +def commSemiRingCatIsoToRingEquiv {R S : CommSemiRingCat.{u}} (e : R ≅ S) : R ≃+* S := + RingEquiv.ofHomInv e.hom.hom e.inv.hom (by ext; simp) (by ext; simp) -end RingEquiv +/-- Build a `RingEquiv` from an isomorphism in the category `CommRingCat`. -/ +def commRingCatIsoToRingEquiv {R S : CommRingCat.{u}} (e : R ≅ S) : R ≃+* S := + RingEquiv.ofHomInv e.hom.hom e.inv.hom (by ext; simp) (by ext; simp) -namespace CategoryTheory.Iso +@[simp] lemma semiRingCatIsoToRingEquiv_toRingHom {R S : SemiRingCat.{u}} (e : R ≅ S) : + (e.semiRingCatIsoToRingEquiv : R →+* S) = e.hom.hom := rfl -/-- Build a `RingEquiv` from an isomorphism in the category `RingCat`. -/ -def ringCatIsoToRingEquiv {X Y : RingCat} (i : X ≅ Y) : X ≃+* Y := - RingEquiv.ofHomInv i.hom i.inv i.hom_inv_id i.inv_hom_id +@[simp] lemma ringCatIsoToRingEquiv_toRingHom {R S : RingCat.{u}} (e : R ≅ S) : + (e.ringCatIsoToRingEquiv : R →+* S) = e.hom.hom := rfl -/-- Build a `RingEquiv` from an isomorphism in the category `CommRingCat`. -/ -def commRingCatIsoToRingEquiv {X Y : CommRingCat} (i : X ≅ Y) : X ≃+* Y := - RingEquiv.ofHomInv i.hom i.inv i.hom_inv_id i.inv_hom_id - --- Porting note: make this high priority to short circuit simplifier -@[simp (high)] -theorem commRingIsoToRingEquiv_toRingHom {X Y : CommRingCat} (i : X ≅ Y) : - i.commRingCatIsoToRingEquiv.toRingHom = i.hom := by - ext - rfl +@[simp] lemma commSemiRingCatIsoToRingEquiv_toRingHom {R S : CommSemiRingCat.{u}} (e : R ≅ S) : + (e.commSemiRingCatIsoToRingEquiv : R →+* S) = e.hom.hom := rfl --- Porting note: make this high priority to short circuit simplifier -@[simp (high)] -theorem commRingIsoToRingEquiv_symm_toRingHom {X Y : CommRingCat} (i : X ≅ Y) : - i.commRingCatIsoToRingEquiv.symm.toRingHom = i.inv := by - ext - rfl +@[simp] lemma commRingCatIsoToRingEquiv_toRingHom {R S : CommRingCat.{u}} (e : R ≅ S) : + (e.commRingCatIsoToRingEquiv : R →+* S) = e.hom.hom := rfl -@[simp] -lemma commRingIsoToRingEquiv_apply {X Y : CommRingCat} (i : X ≅ Y) (x : X) : - i.commRingCatIsoToRingEquiv x = i.hom x := - rfl +end CategoryTheory.Iso -@[simp] -lemma commRingIsoToRingEquiv_symm_apply {X Y : CommRingCat} (i : X ≅ Y) (y : Y) : - i.commRingCatIsoToRingEquiv.symm y = i.inv y := - rfl +-- Porting note: typemax hacks to fix universe complaints +/-- An alias for `SemiringCat.{max u v}`, to deal around unification issues. -/ +@[nolint checkUnivs] +abbrev SemiRingCatMax.{u1, u2} := SemiRingCat.{max u1 u2} -end CategoryTheory.Iso +/-- An alias for `RingCat.{max u v}`, to deal around unification issues. -/ +@[nolint checkUnivs] +abbrev RingCatMax.{u1, u2} := RingCat.{max u1 u2} -/-- Ring equivalences between `RingCat`s are the same as (isomorphic to) isomorphisms in -`RingCat`. -/ -def ringEquivIsoRingIso {X Y : Type u} [Ring X] [Ring Y] : - X ≃+* Y ≅ RingCat.of X ≅ RingCat.of Y where - hom e := e.toRingCatIso - inv i := i.ringCatIsoToRingEquiv - -/-- Ring equivalences between `CommRingCat`s are the same as (isomorphic to) isomorphisms -in `CommRingCat`. -/ -def ringEquivIsoCommRingIso {X Y : Type u} [CommRing X] [CommRing Y] : - X ≃+* Y ≅ CommRingCat.of X ≅ CommRingCat.of Y where - hom e := e.toCommRingCatIso - inv i := i.commRingCatIsoToRingEquiv - -instance RingCat.forget_reflects_isos : (forget RingCat.{u}).ReflectsIsomorphisms where - reflects {X Y} f _ := by - let i := asIso ((forget RingCat).map f) - let ff : X →+* Y := f - let e : X ≃+* Y := { ff, i.toEquiv with } - exact e.toRingCatIso.isIso_hom +/-- An alias for `CommSemiRingCat.{max u v}`, to deal around unification issues. -/ +@[nolint checkUnivs] +abbrev CommSemiRingCatMax.{u1, u2} := CommSemiRingCat.{max u1 u2} -instance CommRingCat.forget_reflects_isos : (forget CommRingCat.{u}).ReflectsIsomorphisms where - reflects {X Y} f _ := by - let i := asIso ((forget CommRingCat).map f) - let ff : X →+* Y := f - let e : X ≃+* Y := { ff, i.toEquiv with } - exact e.toCommRingCatIso.isIso_hom +/-- An alias for `CommRingCat.{max u v}`, to deal around unification issues. -/ +@[nolint checkUnivs] +abbrev CommRingCatMax.{u1, u2} := CommRingCat.{max u1 u2} -theorem CommRingCat.comp_eq_ring_hom_comp {R S T : CommRingCat} (f : R ⟶ S) (g : S ⟶ T) : - f ≫ g = g.comp f := +lemma RingCat.forget_map_apply {R S : RingCat} (f : R ⟶ S) + (x : (CategoryTheory.forget RingCat).obj R) : + @DFunLike.coe _ _ _ ConcreteCategory.instFunLike f x = f x := rfl -theorem CommRingCat.ringHom_comp_eq_comp {R S T : Type _} [CommRing R] [CommRing S] [CommRing T] - (f : R →+* S) (g : S →+* T) : g.comp f = CommRingCat.ofHom f ≫ CommRingCat.ofHom g := +lemma CommRingCat.forget_map_apply {R S : CommRingCat} (f : R ⟶ S) + (x : (CategoryTheory.forget CommRingCat).obj R) : + @DFunLike.coe _ _ _ ConcreteCategory.instFunLike f x = f x := rfl - --- It would be nice if we could have the following, --- but it requires making `reflectsIsomorphisms_forget₂` an instance, --- which can cause typeclass loops: --- Porting note: This was the case in mathlib3, perhaps it is different now? -attribute [local instance] reflectsIsomorphisms_forget₂ - -example : (forget₂ RingCat AddCommGrp).ReflectsIsomorphisms := by infer_instance - -/-! -`@[simp]` lemmas for `RingHom.comp` and categorical identities. --/ - -@[simp] theorem RingHom.comp_id_semiringCat - {G : SemiRingCat.{u}} {H : Type u} [Semiring H] (f : G →+* H) : f.comp (𝟙 G) = f := - Category.id_comp (SemiRingCat.ofHom f) -@[simp] theorem RingHom.id_semiringCat_comp - {G : Type u} [Semiring G] {H : SemiRingCat.{u}} (f : G →+* H) : RingHom.comp (𝟙 H) f = f := - Category.comp_id (SemiRingCat.ofHom f) - -@[simp] theorem RingHom.comp_id_commSemiringCat - {G : CommSemiRingCat.{u}} {H : Type u} [CommSemiring H] (f : G →+* H) : f.comp (𝟙 G) = f := - Category.id_comp (CommSemiRingCat.ofHom f) -@[simp] theorem RingHom.id_commSemiringCat_comp - {G : Type u} [CommSemiring G] {H : CommSemiRingCat.{u}} (f : G →+* H) : - RingHom.comp (𝟙 H) f = f := - Category.comp_id (CommSemiRingCat.ofHom f) - -@[simp] theorem RingHom.comp_id_ringCat - {G : RingCat.{u}} {H : Type u} [Ring H] (f : G →+* H) : f.comp (𝟙 G) = f := - Category.id_comp (RingCat.ofHom f) -@[simp] theorem RingHom.id_ringCat_comp - {G : Type u} [Ring G] {H : RingCat.{u}} (f : G →+* H) : RingHom.comp (𝟙 H) f = f := - Category.comp_id (RingCat.ofHom f) - -@[simp] theorem RingHom.comp_id_commRingCat - {G : CommRingCat.{u}} {H : Type u} [CommRing H] (f : G →+* H) : f.comp (𝟙 G) = f := - Category.id_comp (CommRingCat.ofHom f) -@[simp] theorem RingHom.id_commRingCat_comp - {G : Type u} [CommRing G] {H : CommRingCat.{u}} (f : G →+* H) : RingHom.comp (𝟙 H) f = f := - Category.comp_id (CommRingCat.ofHom f) diff --git a/Mathlib/Algebra/Category/Ring/Colimits.lean b/Mathlib/Algebra/Category/Ring/Colimits.lean index 140d787c20ebf..de2d11a30d083 100644 --- a/Mathlib/Algebra/Category/Ring/Colimits.lean +++ b/Mathlib/Algebra/Category/Ring/Colimits.lean @@ -191,12 +191,12 @@ def coconeFun (j : J) (x : F.obj j) : ColimitType F := /-- The ring homomorphism from a given ring in the diagram to the colimit ring. -/ -def coconeMorphism (j : J) : F.obj j ⟶ colimit F where - toFun := coconeFun F j - map_one' := by apply Quot.sound; apply Relation.one - map_mul' := by intros; apply Quot.sound; apply Relation.mul - map_zero' := by apply Quot.sound; apply Relation.zero - map_add' := by intros; apply Quot.sound; apply Relation.add +def coconeMorphism (j : J) : F.obj j ⟶ colimit F := ofHom + { toFun := coconeFun F j + map_one' := by apply Quot.sound; apply Relation.one + map_mul' := by intros; apply Quot.sound; apply Relation.mul + map_zero' := by apply Quot.sound; apply Relation.zero + map_add' := by intros; apply Quot.sound; apply Relation.add } @[simp] theorem cocone_naturality {j j' : J} (f : j ⟶ j') : @@ -235,7 +235,7 @@ def descFun (s : Cocone F) : ColimitType F → s.pt := by | refl => rfl | symm x y _ ih => exact ih.symm | trans x y z _ _ ih1 ih2 => exact ih1.trans ih2 - | map j j' f x => exact RingHom.congr_fun (s.ι.naturality f) x + | map j j' f x => exact RingHom.congr_fun (congrArg Hom.hom <| s.ι.naturality f) x | zero j => simp | one j => simp | neg j x => simp @@ -261,22 +261,21 @@ def descFun (s : Cocone F) : ColimitType F → s.pt := by /-- The ring homomorphism from the colimit ring to the cone point of any other cocone. -/ -def descMorphism (s : Cocone F) : colimit F ⟶ s.pt where - toFun := descFun F s - map_one' := rfl - map_zero' := rfl - map_add' x y := by - refine Quot.induction_on₂ x y fun a b => ?_ - dsimp [descFun] - rw [← quot_add] - rfl - map_mul' x y := by exact Quot.induction_on₂ x y fun a b => rfl +def descMorphism (s : Cocone F) : colimit F ⟶ s.pt := ofHom + { toFun := descFun F s + map_one' := rfl + map_zero' := rfl + map_add' := fun x y ↦ by + refine Quot.induction_on₂ x y fun a b => ?_ + dsimp [descFun] + rw [← quot_add] + rfl + map_mul' := fun x y ↦ by exact Quot.induction_on₂ x y fun a b => rfl } /-- Evidence that the proposed colimit is the colimit. -/ def colimitIsColimit : IsColimit (colimitCocone F) where desc s := descMorphism F s - uniq s m w := RingHom.ext fun x => by - change (colimitCocone F).pt →+* s.pt at m + uniq s m w := hom_ext <| RingHom.ext fun x => by refine Quot.inductionOn x ?_ intro x induction x with @@ -494,12 +493,12 @@ def coconeFun (j : J) (x : F.obj j) : ColimitType F := /-- The ring homomorphism from a given commutative ring in the diagram to the colimit commutative ring. -/ -def coconeMorphism (j : J) : F.obj j ⟶ colimit F where - toFun := coconeFun F j - map_one' := by apply Quot.sound; apply Relation.one - map_mul' := by intros; apply Quot.sound; apply Relation.mul - map_zero' := by apply Quot.sound; apply Relation.zero - map_add' := by intros; apply Quot.sound; apply Relation.add +def coconeMorphism (j : J) : F.obj j ⟶ colimit F := ofHom <| + { toFun := coconeFun F j + map_one' := by apply Quot.sound; apply Relation.one + map_mul' := by intros; apply Quot.sound; apply Relation.mul + map_zero' := by apply Quot.sound; apply Relation.zero + map_add' := by intros; apply Quot.sound; apply Relation.add } @[simp] theorem cocone_naturality {j j' : J} (f : j ⟶ j') : @@ -538,7 +537,7 @@ def descFun (s : Cocone F) : ColimitType F → s.pt := by | refl => rfl | symm x y _ ih => exact ih.symm | trans x y z _ _ ih1 ih2 => exact ih1.trans ih2 - | map j j' f x => exact RingHom.congr_fun (s.ι.naturality f) x + | map j j' f x => exact RingHom.congr_fun (congrArg Hom.hom <| s.ι.naturality f) x | zero j => simp | one j => simp | neg j x => simp @@ -565,22 +564,21 @@ def descFun (s : Cocone F) : ColimitType F → s.pt := by /-- The ring homomorphism from the colimit commutative ring to the cone point of any other cocone. -/ -def descMorphism (s : Cocone F) : colimit F ⟶ s.pt where - toFun := descFun F s - map_one' := rfl - map_zero' := rfl - map_add' x y := by - refine Quot.induction_on₂ x y fun a b => ?_ - dsimp [descFun] - rw [← quot_add] - rfl - map_mul' x y := by exact Quot.induction_on₂ x y fun a b => rfl +def descMorphism (s : Cocone F) : colimit F ⟶ s.pt := ofHom + { toFun := descFun F s + map_one' := rfl + map_zero' := rfl + map_add' := fun x y ↦ by + refine Quot.induction_on₂ x y fun a b => ?_ + dsimp [descFun] + rw [← quot_add] + rfl + map_mul' := fun x y ↦ by exact Quot.induction_on₂ x y fun a b => rfl } /-- Evidence that the proposed colimit is the colimit. -/ def colimitIsColimit : IsColimit (colimitCocone F) where - desc s := descMorphism F s - uniq s m w := RingHom.ext fun x => by - change (colimitCocone F).pt →+* s.pt at m + desc := fun s ↦ descMorphism F s + uniq := fun s m w ↦ hom_ext <| RingHom.ext fun x => by refine Quot.inductionOn x ?_ intro x induction x with diff --git a/Mathlib/Algebra/Category/Ring/Constructions.lean b/Mathlib/Algebra/Category/Ring/Constructions.lean index 8d91389010d43..dd2e0ccd12b3e 100644 --- a/Mathlib/Algebra/Category/Ring/Constructions.lean +++ b/Mathlib/Algebra/Category/Ring/Constructions.lean @@ -40,8 +40,8 @@ def pushoutCocone : Limits.PushoutCocone (CommRingCat.ofHom (algebraMap R A)) (CommRingCat.ofHom (algebraMap R B)) := by fapply Limits.PushoutCocone.mk · exact CommRingCat.of (A ⊗[R] B) - · exact Algebra.TensorProduct.includeLeftRingHom (A := A) - · exact Algebra.TensorProduct.includeRight.toRingHom (A := B) + · exact ofHom <| Algebra.TensorProduct.includeLeftRingHom (A := A) + · exact ofHom <| Algebra.TensorProduct.includeRight.toRingHom (A := B) · ext r trans algebraMap R (A ⊗[R] B) r · exact Algebra.TensorProduct.includeLeft.commutes (R := R) r @@ -49,12 +49,12 @@ def pushoutCocone : Limits.PushoutCocone @[simp] theorem pushoutCocone_inl : - (pushoutCocone R A B).inl = Algebra.TensorProduct.includeLeftRingHom (A := A) := + (pushoutCocone R A B).inl = ofHom (Algebra.TensorProduct.includeLeftRingHom (A := A)) := rfl @[simp] theorem pushoutCocone_inr : - (pushoutCocone R A B).inr = Algebra.TensorProduct.includeRight.toRingHom (A := B) := + (pushoutCocone R A B).inr = ofHom (Algebra.TensorProduct.includeRight.toRingHom (A := B)) := rfl @[simp] @@ -65,17 +65,18 @@ theorem pushoutCocone_pt : /-- Verify that the `pushout_cocone` is indeed the colimit. -/ def pushoutCoconeIsColimit : Limits.IsColimit (pushoutCocone R A B) := Limits.PushoutCocone.isColimitAux' _ fun s => by - letI := RingHom.toAlgebra (s.inl.comp (algebraMap R A)) + letI := RingHom.toAlgebra (s.inl.hom.comp (algebraMap R A)) let f' : A →ₐ[R] s.pt := - { s.inl with + { s.inl.hom with commutes' := fun r => rfl } let g' : B →ₐ[R] s.pt := - { s.inr with - commutes' := DFunLike.congr_fun ((s.ι.naturality Limits.WalkingSpan.Hom.snd).trans - (s.ι.naturality Limits.WalkingSpan.Hom.fst).symm) } + { s.inr.hom with + commutes' := DFunLike.congr_fun <| congrArg Hom.hom + ((s.ι.naturality Limits.WalkingSpan.Hom.snd).trans + (s.ι.naturality Limits.WalkingSpan.Hom.fst).symm) } letI : Algebra R (pushoutCocone R A B).pt := show Algebra R (A ⊗[R] B) by infer_instance -- The factor map is a ⊗ b ↦ f(a) * g(b). - use AlgHom.toRingHom (Algebra.TensorProduct.productMap f' g') + use ofHom (AlgHom.toRingHom (Algebra.TensorProduct.productMap f' g')) simp only [pushoutCocone_inl, pushoutCocone_inr] constructor · ext x @@ -85,7 +86,7 @@ def pushoutCoconeIsColimit : Limits.IsColimit (pushoutCocone R A B) := exact Algebra.TensorProduct.productMap_right_apply (B := B) _ _ x intro h eq1 eq2 let h' : A ⊗[R] B →ₐ[R] s.pt := - { h with + { h.hom with commutes' := fun r => by change h (algebraMap R A r ⊗ₜ[R] 1) = s.inl (algebraMap R A r) rw [← eq1] @@ -100,7 +101,7 @@ def pushoutCoconeIsColimit : Limits.IsColimit (pushoutCocone R A B) := simp only [f', g', ← eq1, pushoutCocone_pt, ← eq2, AlgHom.toRingHom_eq_coe, Algebra.TensorProduct.productMap_apply_tmul, AlgHom.coe_mk] change _ = h (a ⊗ₜ 1) * h (1 ⊗ₜ b) - rw [← h.map_mul, Algebra.TensorProduct.tmul_mul_tmul, mul_one, one_mul] + rw [← h.hom.map_mul, Algebra.TensorProduct.tmul_mul_tmul, mul_one, one_mul] rfl lemma isPushout_tensorProduct (R A B : Type u) [CommRing R] [CommRing A] [CommRing B] @@ -117,24 +118,24 @@ end Pushout section Terminal +instance (X : CommRingCat.{u}) : Unique (X ⟶ CommRingCat.of.{u} PUnit) := + ⟨⟨ofHom <| ⟨1, rfl, by simp⟩⟩, fun f ↦ by ext⟩ + /-- The trivial ring is the (strict) terminal object of `CommRingCat`. -/ -def punitIsTerminal : IsTerminal (CommRingCat.of.{u} PUnit) := by - refine IsTerminal.ofUnique (h := fun X => ⟨⟨⟨⟨1, rfl⟩, fun _ _ => rfl⟩, ?_, ?_⟩, ?_⟩) - · rfl - · intros; simp only [coe_of, Pi.one_apply, self_eq_add_right]; ext - · intros f; ext; rfl +def punitIsTerminal : IsTerminal (CommRingCat.of.{u} PUnit) := + IsTerminal.ofUnique _ instance commRingCat_hasStrictTerminalObjects : HasStrictTerminalObjects CommRingCat.{u} := by apply hasStrictTerminalObjects_of_terminal_is_strict (CommRingCat.of PUnit) intro X f - refine ⟨⟨⟨1, rfl, fun _ _ => rfl⟩, by ext; rfl, ?_⟩⟩ - ext x - have e : (0 : X) = 1 := by - rw [← f.map_one, ← f.map_zero] - congr - replace e : 0 * x = 1 * x := congr_arg (· * x) e - rw [one_mul, zero_mul, ← f.map_zero] at e - exact e + refine ⟨ofHom ⟨1, rfl, by simp⟩, ?_, ?_⟩ + · ext + · ext x + have e : (0 : X) = 1 := by + rw [← f.hom.map_one, ← f.hom.map_zero] + replace e : 0 * x = 1 * x := congr_arg (· * x) e + rw [one_mul, zero_mul, ← f.hom.map_zero] at e + exact e theorem subsingleton_of_isTerminal {X : CommRingCat} (hX : IsTerminal X) : Subsingleton X := (hX.uniqueUpToIso punitIsTerminal).commRingCatIsoToRingEquiv.toEquiv.subsingleton_congr.mpr @@ -142,12 +143,14 @@ theorem subsingleton_of_isTerminal {X : CommRingCat} (hX : IsTerminal X) : Subsi /-- `ℤ` is the initial object of `CommRingCat`. -/ def zIsInitial : IsInitial (CommRingCat.of ℤ) := - IsInitial.ofUnique (h := fun R => ⟨⟨Int.castRingHom R⟩, fun a => a.ext_int _⟩) + IsInitial.ofUnique (h := fun R => ⟨⟨ofHom <| Int.castRingHom R⟩, + fun a => hom_ext <| a.hom.ext_int _⟩) /-- `ULift.{u} ℤ` is initial in `CommRingCat`. -/ def isInitial : IsInitial (CommRingCat.of (ULift.{u} ℤ)) := - IsInitial.ofUnique (h := fun R ↦ ⟨⟨(Int.castRingHom R).comp ULift.ringEquiv.toRingHom⟩, + IsInitial.ofUnique (h := fun R ↦ ⟨⟨ofHom <| (Int.castRingHom R).comp ULift.ringEquiv.toRingHom⟩, fun _ ↦ by + ext : 1 rw [← RingHom.cancel_right (f := (ULift.ringEquiv.{0, u} (α := ℤ)).symm.toRingHom) (hf := ULift.ringEquiv.symm.surjective)] apply RingHom.ext_int⟩) @@ -165,7 +168,7 @@ def prodFan : BinaryFan A B := /-- The product in `CommRingCat` is the cartesian product. -/ def prodFanIsLimit : IsLimit (prodFan A B) where - lift c := RingHom.prod (c.π.app ⟨WalkingPair.left⟩) (c.π.app ⟨WalkingPair.right⟩) + lift c := ofHom <| RingHom.prod (c.π.app ⟨WalkingPair.left⟩).hom (c.π.app ⟨WalkingPair.right⟩).hom fac c j := by ext rcases j with ⟨⟨⟩⟩ <;> @@ -175,12 +178,12 @@ def prodFanIsLimit : IsLimit (prodFan A B) where uniq s m h := by ext x change m x = (BinaryFan.fst s x, BinaryFan.snd s x) - have eq1 := congr_hom (h ⟨WalkingPair.left⟩) x - have eq2 := congr_hom (h ⟨WalkingPair.right⟩) x - dsimp at eq1 eq2 - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [← eq1, ← eq2] - rfl + have eq1 : (m ≫ (A.prodFan B).fst) x = (BinaryFan.fst s) x := + congr_hom (h ⟨WalkingPair.left⟩) x + have eq2 : (m ≫ (A.prodFan B).snd) x = (BinaryFan.snd s) x := + congr_hom (h ⟨WalkingPair.right⟩) x + rw [← eq1, ← eq2] + simp [prodFan] end Product @@ -193,15 +196,16 @@ The categorical product of rings is the cartesian product of rings. This is its -/ @[simps! pt] def piFan : Fan R := - Fan.mk (CommRingCat.of ((i : ι) → R i)) (Pi.evalRingHom _) + Fan.mk (CommRingCat.of ((i : ι) → R i)) (fun i ↦ ofHom <| Pi.evalRingHom _ i) /-- The categorical product of rings is the cartesian product of rings. -/ def piFanIsLimit : IsLimit (piFan R) where - lift s := Pi.ringHom fun i ↦ s.π.1 ⟨i⟩ + lift s := ofHom <| Pi.ringHom fun i ↦ (s.π.1 ⟨i⟩).hom fac s i := by rfl - uniq _ _ h := DFunLike.ext _ _ fun x ↦ funext fun i ↦ DFunLike.congr_fun (h ⟨i⟩) x + uniq _ _ h := hom_ext <| DFunLike.ext _ _ fun x ↦ funext fun i ↦ + DFunLike.congr_fun (congrArg Hom.hom <| h ⟨i⟩) x /-- The categorical product and the usual product agrees @@ -224,7 +228,7 @@ variable {A B : CommRingCat.{u}} (f g : A ⟶ B) /-- The equalizer in `CommRingCat` is the equalizer as sets. This is the equalizer fork. -/ def equalizerFork : Fork f g := - Fork.ofι (CommRingCat.ofHom (RingHom.eqLocus f g).subtype) <| by + Fork.ofι (CommRingCat.ofHom (RingHom.eqLocus f.hom g.hom).subtype) <| by ext ⟨x, e⟩ simpa using e @@ -232,29 +236,27 @@ def equalizerFork : Fork f g := def equalizerForkIsLimit : IsLimit (equalizerFork f g) := by fapply Fork.IsLimit.mk' intro s - -- Porting note: Lean can't see through `(parallelPair f g).obj zero` - haveI : SubsemiringClass (Subring A) ((parallelPair f g).obj WalkingParallelPair.zero) := - show SubsemiringClass (Subring A) A by infer_instance - use s.ι.codRestrict _ fun x => (ConcreteCategory.congr_hom s.condition x : _) + use ofHom <| s.ι.hom.codRestrict _ fun x => (ConcreteCategory.congr_hom s.condition x : _) constructor · ext rfl · intro m hm - exact RingHom.ext fun x => Subtype.ext <| ConcreteCategory.congr_hom hm x + ext x + exact Subtype.ext <| RingHom.congr_fun (congrArg Hom.hom hm) x -instance : IsLocalHom (equalizerFork f g).ι := by +instance : IsLocalHom (equalizerFork f g).ι.hom := by constructor rintro ⟨a, h₁ : _ = _⟩ (⟨⟨x, y, h₃, h₄⟩, rfl : x = _⟩ : IsUnit a) - have : y ∈ RingHom.eqLocus f g := by - apply (f.isUnit_map ⟨⟨x, y, h₃, h₄⟩, rfl⟩ : IsUnit (f x)).mul_left_inj.mp + have : y ∈ RingHom.eqLocus f.hom g.hom := by + apply (f.hom.isUnit_map ⟨⟨x, y, h₃, h₄⟩, rfl⟩ : IsUnit (f x)).mul_left_inj.mp conv_rhs => rw [h₁] - rw [← f.map_mul, ← g.map_mul, h₄, f.map_one, g.map_one] + rw [← f.hom.map_mul, ← g.hom.map_mul, h₄, f.hom.map_one, g.hom.map_one] rw [isUnit_iff_exists_inv] exact ⟨⟨y, this⟩, Subtype.eq h₃⟩ @[instance] theorem equalizer_ι_isLocalHom (F : WalkingParallelPair ⥤ CommRingCat.{u}) : - IsLocalHom (limit.π F WalkingParallelPair.zero) := by + IsLocalHom (limit.π F WalkingParallelPair.zero).hom := by have := limMap_π (diagramIsoParallelPair F).hom WalkingParallelPair.zero rw [← IsIso.comp_inv_eq] at this rw [← this] @@ -263,7 +265,7 @@ theorem equalizer_ι_isLocalHom (F : WalkingParallelPair ⥤ CommRingCat.{u}) : equalizerForkIsLimit (F.map WalkingParallelPairHom.left) (F.map WalkingParallelPairHom.right)⟩ WalkingParallelPair.zero] - change IsLocalHom ((lim.map _ ≫ _ ≫ (equalizerFork _ _).ι) ≫ _) + change IsLocalHom ((lim.map _ ≫ _ ≫ (equalizerFork _ _).ι) ≫ _).hom infer_instance @[deprecated (since := "2024-10-10")] @@ -274,12 +276,15 @@ open CategoryTheory.Limits.WalkingParallelPair Opposite open CategoryTheory.Limits.WalkingParallelPairHom instance equalizer_ι_is_local_ring_hom' (F : WalkingParallelPairᵒᵖ ⥤ CommRingCat.{u}) : - IsLocalHom (limit.π F (Opposite.op WalkingParallelPair.one)) := by + IsLocalHom (limit.π F (Opposite.op WalkingParallelPair.one)).hom := by have : _ = limit.π F (walkingParallelPairOpEquiv.functor.obj _) := (limit.isoLimitCone_inv_π ⟨_, IsLimit.whiskerEquivalence (limit.isLimit F) walkingParallelPairOpEquiv⟩ WalkingParallelPair.zero : _) erw [← this] + -- note: this was not needed before https://github.com/leanprover-community/mathlib4/pull/19757 + haveI : IsLocalHom (limit.π (walkingParallelPairOpEquiv.functor ⋙ F) zero).hom := by + infer_instance infer_instance end Equalizer @@ -293,10 +298,10 @@ def pullbackCone {A B C : CommRingCat.{u}} (f : A ⟶ C) (g : B ⟶ C) : Pullbac PullbackCone.mk (CommRingCat.ofHom <| (RingHom.fst A B).comp - (RingHom.eqLocus (f.comp (RingHom.fst A B)) (g.comp (RingHom.snd A B))).subtype) + (RingHom.eqLocus (f.hom.comp (RingHom.fst A B)) (g.hom.comp (RingHom.snd A B))).subtype) (CommRingCat.ofHom <| (RingHom.snd A B).comp - (RingHom.eqLocus (f.comp (RingHom.fst A B)) (g.comp (RingHom.snd A B))).subtype) + (RingHom.eqLocus (f.hom.comp (RingHom.fst A B)) (g.hom.comp (RingHom.snd A B))).subtype) (by ext ⟨x, e⟩ simpa [CommRingCat.ofHom] using e) @@ -306,9 +311,9 @@ def pullbackConeIsLimit {A B C : CommRingCat.{u}} (f : A ⟶ C) (g : B ⟶ C) : IsLimit (pullbackCone f g) := by fapply PullbackCone.IsLimit.mk · intro s - apply (s.fst.prod s.snd).codRestrict + refine ofHom ((s.fst.hom.prod s.snd.hom).codRestrict _ ?_) intro x - exact congr_arg (fun f : s.pt →+* C => f x) s.condition + exact congr_arg (fun f : s.pt →+* C => f x) (congrArg Hom.hom s.condition) · intro s ext x rfl @@ -316,10 +321,10 @@ def pullbackConeIsLimit {A B C : CommRingCat.{u}} (f : A ⟶ C) (g : B ⟶ C) : ext x rfl · intro s m e₁ e₂ - refine RingHom.ext fun (x : s.pt) => Subtype.ext ?_ + refine hom_ext <| RingHom.ext fun (x : s.pt) => Subtype.ext ?_ change (m x).1 = (_, _) - have eq1 := (congr_arg (fun f : s.pt →+* A => f x) e₁ : _) - have eq2 := (congr_arg (fun f : s.pt →+* B => f x) e₂ : _) + have eq1 := (congr_arg (fun f : s.pt →+* A => f x) (congrArg Hom.hom e₁) : _) + have eq2 := (congr_arg (fun f : s.pt →+* B => f x) (congrArg Hom.hom e₂) : _) rw [← eq1, ← eq2] rfl diff --git a/Mathlib/Algebra/Category/Ring/Epi.lean b/Mathlib/Algebra/Category/Ring/Epi.lean index a961e33cdf72e..ae3808ac12470 100644 --- a/Mathlib/Algebra/Category/Ring/Epi.lean +++ b/Mathlib/Algebra/Category/Ring/Epi.lean @@ -30,22 +30,21 @@ lemma CommRingCat.epi_iff_tmul_eq_tmul {R S : Type u} [CommRing R] [CommRing S] (CommRingCat.ofHom <| (Algebra.TensorProduct.includeRight (R := R) (A := S)).toRingHom) (by ext r; show algebraMap R S r ⊗ₜ 1 = 1 ⊗ₜ algebraMap R S r; simp only [Algebra.algebraMap_eq_smul_one, smul_tmul]) - exact RingHom.congr_fun this + exact RingHom.congr_fun (congrArg Hom.hom this) · refine fun H ↦ ⟨fun {T} f g e ↦ ?_⟩ - letI : Algebra R T := (ofHom (algebraMap R S) ≫ g).toAlgebra - let f' : S →ₐ[R] T := ⟨f, RingHom.congr_fun e⟩ - let g' : S →ₐ[R] T := ⟨g, fun _ ↦ rfl⟩ + letI : Algebra R T := (ofHom (algebraMap R S) ≫ g).hom.toAlgebra + let f' : S →ₐ[R] T := ⟨f.hom, RingHom.congr_fun (congrArg Hom.hom e)⟩ + let g' : S →ₐ[R] T := ⟨g.hom, fun _ ↦ rfl⟩ ext s simpa using congr(Algebra.TensorProduct.lift f' g' (fun _ _ ↦ .all _ _) $(H s)) lemma RingHom.surjective_of_epi_of_finite {R S : CommRingCat} (f : R ⟶ S) [Epi f] - (h₂ : RingHom.Finite f) : Function.Surjective f := by - letI := f.toAlgebra - have : Module.Finite R S := h₂ + (h₂ : RingHom.Finite f.hom) : Function.Surjective f := by + algebraize [f.hom] apply RingHom.surjective_of_tmul_eq_tmul_of_finite rwa [← CommRingCat.epi_iff_tmul_eq_tmul] lemma RingHom.surjective_iff_epi_and_finite {R S : CommRingCat} {f : R ⟶ S} : - Function.Surjective f ↔ Epi f ∧ RingHom.Finite f where - mp h := ⟨ConcreteCategory.epi_of_surjective f h, .of_surjective f h⟩ + Function.Surjective f ↔ Epi f ∧ RingHom.Finite f.hom where + mp h := ⟨ConcreteCategory.epi_of_surjective f h, .of_surjective f.hom h⟩ mpr := fun ⟨_, h⟩ ↦ surjective_of_epi_of_finite f h diff --git a/Mathlib/Algebra/Category/Ring/FilteredColimits.lean b/Mathlib/Algebra/Category/Ring/FilteredColimits.lean index 197ff8afa8f5c..1f10eb428fe0c 100644 --- a/Mathlib/Algebra/Category/Ring/FilteredColimits.lean +++ b/Mathlib/Algebra/Category/Ring/FilteredColimits.lean @@ -111,31 +111,32 @@ def colimit : SemiRingCatMax.{v, u} := def colimitCocone : Cocone F where pt := colimit.{v, u} F ι := - { app := fun j => + { app := fun j => ofHom { (MonCat.FilteredColimits.colimitCocone (F ⋙ forget₂ SemiRingCatMax.{v, u} MonCat)).ι.app j, (AddCommMonCat.FilteredColimits.colimitCocone (F ⋙ forget₂ SemiRingCatMax.{v, u} AddCommMonCat)).ι.app j with } - naturality := fun {_ _} f => + naturality := fun {_ _} f => hom_ext <| RingHom.coe_inj ((Types.TypeMax.colimitCocone (F ⋙ forget SemiRingCat)).ι.naturality f) } /-- The proposed colimit cocone is a colimit in `SemiRingCat`. -/ def colimitCoconeIsColimit : IsColimit <| colimitCocone.{v, u} F where - desc t := + desc t := ofHom { (MonCat.FilteredColimits.colimitCoconeIsColimit.{v, u} (F ⋙ forget₂ SemiRingCatMax.{v, u} MonCat)).desc ((forget₂ SemiRingCatMax.{v, u} MonCat).mapCocone t), (AddCommMonCat.FilteredColimits.colimitCoconeIsColimit.{v, u} (F ⋙ forget₂ SemiRingCatMax.{v, u} AddCommMonCat)).desc ((forget₂ SemiRingCatMax.{v, u} AddCommMonCat).mapCocone t) with } - fac t j := + fac t j := hom_ext <| RingHom.coe_inj <| (Types.TypeMax.colimitCoconeIsColimit.{v, u} (F ⋙ forget SemiRingCatMax.{v, u})).fac ((forget SemiRingCatMax.{v, u}).mapCocone t) j - uniq t _ h := + uniq t _ h := hom_ext <| RingHom.coe_inj <| (Types.TypeMax.colimitCoconeIsColimit (F ⋙ forget SemiRingCat)).uniq - ((forget SemiRingCat).mapCocone t) _ fun j => funext fun x => RingHom.congr_fun (h j) x + ((forget SemiRingCat).mapCocone t) _ fun j => funext fun x => + RingHom.congr_fun (congrArg Hom.hom (h j)) x instance forget₂Mon_preservesFilteredColimits : PreservesFilteredColimits (forget₂ SemiRingCat MonCat.{u}) where @@ -167,7 +168,7 @@ abbrev R : SemiRingCatMax.{v, u} := SemiRingCat.FilteredColimits.colimit (F ⋙ forget₂ CommSemiRingCat SemiRingCat.{max v u}) instance colimitCommSemiring : CommSemiring.{max v u} <| R.{v, u} F := - { (R F).str, + { (R F).semiring, CommMonCat.FilteredColimits.colimitCommMonoid (F ⋙ forget₂ CommSemiRingCat CommMonCat.{max v u}) with } @@ -179,23 +180,27 @@ def colimit : CommSemiRingCat.{max v u} := def colimitCocone : Cocone F where pt := colimit.{v, u} F ι := - { (SemiRingCat.FilteredColimits.colimitCocone - (F ⋙ forget₂ CommSemiRingCat SemiRingCat.{max v u})).ι with } + { app := fun X ↦ ofHom <| ((SemiRingCat.FilteredColimits.colimitCocone + (F ⋙ forget₂ CommSemiRingCat SemiRingCat.{max v u})).ι.app X).hom + naturality := fun _ _ f ↦ hom_ext <| + RingHom.coe_inj ((Types.TypeMax.colimitCocone + (F ⋙ forget CommSemiRingCat)).ι.naturality f) } /-- The proposed colimit cocone is a colimit in `CommSemiRingCat`. -/ def colimitCoconeIsColimit : IsColimit <| colimitCocone.{v, u} F where - desc t := + desc t := ofHom <| (SemiRingCat.FilteredColimits.colimitCoconeIsColimit.{v, u} (F ⋙ forget₂ CommSemiRingCat SemiRingCat.{max v u})).desc - ((forget₂ CommSemiRingCat SemiRingCat).mapCocone t) - fac t j := + ((forget₂ CommSemiRingCat SemiRingCat).mapCocone t) |>.hom + fac t j := hom_ext <| RingHom.coe_inj <| (Types.TypeMax.colimitCoconeIsColimit.{v, u} (F ⋙ forget CommSemiRingCat)).fac ((forget CommSemiRingCat).mapCocone t) j - uniq t _ h := + uniq t _ h := hom_ext <| RingHom.coe_inj <| (Types.TypeMax.colimitCoconeIsColimit (F ⋙ forget CommSemiRingCat)).uniq - ((forget CommSemiRingCat).mapCocone t) _ fun j => funext fun x => RingHom.congr_fun (h j) x + ((forget CommSemiRingCat).mapCocone t) _ fun j => funext fun x => + RingHom.congr_fun (congrArg Hom.hom (h j)) x instance forget₂SemiRing_preservesFilteredColimits : PreservesFilteredColimits (forget₂ CommSemiRingCat SemiRingCat.{u}) where @@ -230,7 +235,7 @@ abbrev R : SemiRingCat.{max v u} := SemiRingCat.FilteredColimits.colimit.{v, u} (F ⋙ forget₂ RingCat SemiRingCat.{max v u}) instance colimitRing : Ring.{max v u} <| R.{v, u} F := - { (R F).str, + { (R F).semiring, AddCommGrp.FilteredColimits.colimitAddCommGroup.{v, u} (F ⋙ forget₂ RingCat AddCommGrp.{max v u}) with } @@ -242,23 +247,26 @@ def colimit : RingCat.{max v u} := def colimitCocone : Cocone F where pt := colimit.{v, u} F ι := - { (SemiRingCat.FilteredColimits.colimitCocone - (F ⋙ forget₂ RingCat SemiRingCat.{max v u})).ι with } + { app := fun X ↦ ofHom <| ((SemiRingCat.FilteredColimits.colimitCocone + (F ⋙ forget₂ RingCat SemiRingCat.{max v u})).ι.app X).hom + naturality := fun _ _ f ↦ hom_ext <| + RingHom.coe_inj ((Types.TypeMax.colimitCocone (F ⋙ forget RingCat)).ι.naturality f) } /-- The proposed colimit cocone is a colimit in `Ring`. -/ def colimitCoconeIsColimit : IsColimit <| colimitCocone.{v, u} F where - desc t := + desc t := ofHom <| (SemiRingCat.FilteredColimits.colimitCoconeIsColimit.{v, u} (F ⋙ forget₂ RingCat SemiRingCat.{max v u})).desc - ((forget₂ RingCat SemiRingCat).mapCocone t) - fac t j := + ((forget₂ RingCat SemiRingCat).mapCocone t) |>.hom + fac t j := hom_ext <| RingHom.coe_inj <| (Types.TypeMax.colimitCoconeIsColimit.{v, u} (F ⋙ forget RingCat)).fac ((forget RingCat).mapCocone t) j - uniq t _ h := + uniq t _ h := hom_ext <| RingHom.coe_inj <| (Types.TypeMax.colimitCoconeIsColimit (F ⋙ forget RingCat)).uniq - ((forget RingCat).mapCocone t) _ fun j => funext fun x => RingHom.congr_fun (h j) x + ((forget RingCat).mapCocone t) _ fun j => funext fun x => + RingHom.congr_fun (congrArg Hom.hom (h j)) x instance forget₂SemiRing_preservesFilteredColimits : PreservesFilteredColimits (forget₂ RingCat SemiRingCat.{u}) where @@ -291,7 +299,7 @@ abbrev R : RingCat.{max v u} := RingCat.FilteredColimits.colimit.{v, u} (F ⋙ forget₂ CommRingCat RingCat.{max v u}) instance colimitCommRing : CommRing.{max v u} <| R.{v, u} F := - { (R.{v, u} F).str, + { (R.{v, u} F).ring, CommSemiRingCat.FilteredColimits.colimitCommSemiring (F ⋙ forget₂ CommRingCat CommSemiRingCat.{max v u}) with } @@ -303,22 +311,26 @@ def colimit : CommRingCat.{max v u} := def colimitCocone : Cocone F where pt := colimit.{v, u} F ι := - { (RingCat.FilteredColimits.colimitCocone (F ⋙ forget₂ CommRingCat RingCat.{max v u})).ι with } + { app := fun X ↦ ofHom <| ((RingCat.FilteredColimits.colimitCocone + (F ⋙ forget₂ CommRingCat RingCat.{max v u})).ι.app X).hom + naturality := fun _ _ f ↦ hom_ext <| + RingHom.coe_inj ((Types.TypeMax.colimitCocone (F ⋙ forget CommRingCat)).ι.naturality f) } /-- The proposed colimit cocone is a colimit in `CommRingCat`. -/ def colimitCoconeIsColimit : IsColimit <| colimitCocone.{v, u} F where - desc t := + desc t := ofHom <| (RingCat.FilteredColimits.colimitCoconeIsColimit.{v, u} (F ⋙ forget₂ CommRingCat RingCat.{max v u})).desc - ((forget₂ CommRingCat RingCat).mapCocone t) - fac t j := + ((forget₂ CommRingCat RingCat).mapCocone t) |>.hom + fac t j := hom_ext <| RingHom.coe_inj <| (Types.TypeMax.colimitCoconeIsColimit.{v, u} (F ⋙ forget CommRingCat)).fac ((forget CommRingCat).mapCocone t) j - uniq t _ h := + uniq t _ h := hom_ext <| RingHom.coe_inj <| (Types.TypeMax.colimitCoconeIsColimit (F ⋙ forget CommRingCat)).uniq - ((forget CommRingCat).mapCocone t) _ fun j => funext fun x => RingHom.congr_fun (h j) x + ((forget CommRingCat).mapCocone t) _ fun j => funext fun x => + RingHom.congr_fun (congrArg Hom.hom <| h j) x instance forget₂Ring_preservesFilteredColimits : PreservesFilteredColimits (forget₂ CommRingCat RingCat.{u}) where diff --git a/Mathlib/Algebra/Category/Ring/Instances.lean b/Mathlib/Algebra/Category/Ring/Instances.lean index def5d4f5edf11..4fcf3b005de52 100644 --- a/Mathlib/Algebra/Category/Ring/Instances.lean +++ b/Mathlib/Algebra/Category/Ring/Instances.lean @@ -26,7 +26,8 @@ instance localization_unit_isIso' (R : CommRingCat) : theorem IsLocalization.epi {R : Type*} [CommRing R] (M : Submonoid R) (S : Type _) [CommRing S] [Algebra R S] [IsLocalization M S] : Epi (CommRingCat.ofHom <| algebraMap R S) := - ⟨fun {T} _ _ => @IsLocalization.ringHom_ext R _ M S _ _ T _ _ _ _⟩ + ⟨fun {T} _ _ h => CommRingCat.hom_ext <| + @IsLocalization.ringHom_ext R _ M S _ _ T _ _ _ _ (congrArg CommRingCat.Hom.hom h)⟩ instance Localization.epi {R : Type*} [CommRing R] (M : Submonoid R) : Epi (CommRingCat.ofHom <| algebraMap R <| Localization M) := @@ -37,44 +38,18 @@ instance Localization.epi' {R : CommRingCat} (M : Submonoid R) : rcases R with ⟨α, str⟩ exact IsLocalization.epi M _ -/- -These three instances solve the issue where the `FunLike` instances provided by -`CommRingCat.instFunLike'`, `CommRingCat.instFunLike''`, and `CommRingCat.instFunLike'''` -are not syntactically equal to `CommRingCat.instFunLike` when applied to -objects of the form `CommRingCat.of R`. -To prevent infinite loops, the priority of these three instances must be set lower -than that of other instances. --/ -instance (priority := 50) {R : Type*} [CommRing R] {S : CommRingCat} (f : CommRingCat.of R ⟶ S) - [IsLocalHom (R := CommRingCat.of R) f] : IsLocalHom f := - inferInstance - -instance (priority := 50) {R : CommRingCat} {S : Type*} [CommRing S] (f : R ⟶ CommRingCat.of S) - [IsLocalHom (S := CommRingCat.of S) f] : IsLocalHom f := - inferInstance - -instance (priority := 50) {R S : Type u} [CommRing R] [CommRing S] - (f : CommRingCat.of R ⟶ CommRingCat.of S) - [IsLocalHom (R := CommRingCat.of R) (S := CommRingCat.of S) f] : IsLocalHom f := - inferInstance - --- This instance handles the coercion of a morphism into a real `RingHom`. -instance {R S : CommRingCat} (f : R ⟶ S) [IsLocalHom f] : - IsLocalHom (F := R →+* S) f := - inferInstance - @[instance] theorem CommRingCat.isLocalHom_comp {R S T : CommRingCat} (f : R ⟶ S) (g : S ⟶ T) - [IsLocalHom g] [IsLocalHom f] : IsLocalHom (f ≫ g) := + [IsLocalHom g.hom] [IsLocalHom f.hom] : IsLocalHom (f ≫ g).hom := RingHom.isLocalHom_comp _ _ @[deprecated (since := "2024-10-10")] alias CommRingCat.isLocalRingHom_comp := CommRingCat.isLocalHom_comp -theorem isLocalHom_of_iso {R S : CommRingCat} (f : R ≅ S) : IsLocalHom f.hom := +theorem isLocalHom_of_iso {R S : CommRingCat} (f : R ≅ S) : IsLocalHom f.hom.hom := { map_nonunit := fun a ha => by - convert f.inv.isUnit_map ha - exact (RingHom.congr_fun f.hom_inv_id _).symm } + convert f.inv.hom.isUnit_map ha + simp } @[deprecated (since := "2024-10-10")] alias isLocalRingHom_of_iso := isLocalHom_of_iso @@ -82,7 +57,7 @@ alias isLocalRingHom_of_iso := isLocalHom_of_iso -- see Note [lower instance priority] @[instance 100] theorem isLocalHom_of_isIso {R S : CommRingCat} (f : R ⟶ S) [IsIso f] : - IsLocalHom f := + IsLocalHom f.hom := isLocalHom_of_iso (asIso f) @[deprecated (since := "2024-10-10")] diff --git a/Mathlib/Algebra/Category/Ring/Limits.lean b/Mathlib/Algebra/Category/Ring/Limits.lean index 7a46166affcf1..a915fc00eabac 100644 --- a/Mathlib/Algebra/Category/Ring/Limits.lean +++ b/Mathlib/Algebra/Category/Ring/Limits.lean @@ -85,8 +85,8 @@ namespace HasLimits def limitCone : Cone F where pt := SemiRingCat.of (Types.Small.limitCone (F ⋙ forget _)).pt π := - { app := limitπRingHom.{v, u} F - naturality := fun {_ _} f => RingHom.coe_inj + { app := fun j ↦ SemiRingCat.ofHom <| limitπRingHom.{v, u} F j + naturality := fun {_ _} f ↦ hom_ext <| RingHom.coe_inj ((Types.Small.limitCone (F ⋙ forget _)).π.naturality f) } /-- Witness that the limit cone in `SemiRingCat` is a limit cone. @@ -94,7 +94,7 @@ def limitCone : Cone F where -/ def limitConeIsLimit : IsLimit (limitCone F) := by refine IsLimit.ofFaithful (forget SemiRingCat.{u}) (Types.Small.limitConeIsLimit.{v, u} _) - (fun s => { toFun := _, map_one' := ?_, map_mul' := ?_, map_zero' := ?_, map_add' := ?_}) + (fun s => ofHom { toFun := _, map_one' := ?_, map_mul' := ?_, map_zero' := ?_, map_add' := ?_}) (fun s => rfl) · simp only [Functor.mapCone_π_app, forget_map, map_one] rfl @@ -184,11 +184,6 @@ instance forget_preservesLimits : PreservesLimits (forget SemiRingCat.{u}) := end SemiRingCat --- Porting note: typemax hack to fix universe complaints -/-- An alias for `CommSemiring.{max u v}`, to deal with unification issues. -/ -@[nolint checkUnivs] -abbrev CommSemiRingCatMax.{u1, u2} := CommSemiRingCat.{max u1 u2} - namespace CommSemiRingCat variable {J : Type v} [Category.{w} J] (F : J ⥤ CommSemiRingCat.{u}) @@ -228,8 +223,9 @@ instance : π := { app := fun j => CommSemiRingCat.ofHom <| SemiRingCat.limitπRingHom.{v, u} (J := J) (F ⋙ forget₂ CommSemiRingCat.{u} SemiRingCat.{u}) j - naturality := (SemiRingCat.HasLimits.limitCone.{v, u} - (F ⋙ forget₂ CommSemiRingCat.{u} SemiRingCat.{u})).π.naturality } } + naturality := fun _ _ f ↦ hom_ext <| congrArg SemiRingCat.Hom.hom <| + (SemiRingCat.HasLimits.limitCone.{v, u} + (F ⋙ forget₂ CommSemiRingCat.{u} SemiRingCat.{u})).π.naturality f } } createsLimitOfReflectsIso fun c' t => { liftedCone := c validLift := IsLimit.uniqueUpToIso (SemiRingCat.HasLimits.limitConeIsLimit.{v, u} _) t @@ -291,11 +287,6 @@ instance forget_preservesLimits : PreservesLimits (forget CommSemiRingCat.{u}) : end CommSemiRingCat --- Porting note: typemax hack to fix universe complaints -/-- An alias for `RingCat.{max u v}`, to deal around unification issues. -/ -@[nolint checkUnivs] -abbrev RingCatMax.{u1, u2} := RingCat.{max u1 u2} - namespace RingCat variable {J : Type v} [Category.{w} J] (F : J ⥤ RingCat.{u}) @@ -334,7 +325,7 @@ instance : CreatesLimit F (forget₂ RingCat.{u} SemiRingCat.{u}) := { pt := RingCat.of (Types.Small.limitCone (F ⋙ forget _)).pt π := { app := fun x => ofHom <| SemiRingCat.limitπRingHom.{v, u} (F ⋙ forget₂ _ SemiRingCat) x - naturality := fun _ _ f => RingHom.coe_inj + naturality := fun _ _ f => hom_ext <| RingHom.coe_inj ((Types.Small.limitCone (F ⋙ forget _)).π.naturality f) } } createsLimitOfReflectsIso fun c' t => { liftedCone := c @@ -422,11 +413,6 @@ instance forget_preservesLimits : PreservesLimits (forget RingCat.{u}) := end RingCat --- Porting note: typemax hack to fix universe complaints -/-- An alias for `CommRingCat.{max u v}`, to deal around unification issues. -/ -@[nolint checkUnivs] -abbrev CommRingCatMax.{u1, u2} := CommRingCat.{max u1 u2} - namespace CommRingCat variable {J : Type v} [Category.{w} J] (F : J ⥤ CommRingCat.{u}) @@ -469,7 +455,7 @@ instance : π := { app := fun x => ofHom <| SemiRingCat.limitπRingHom.{v, u} F' x naturality := - fun _ _ f => RingHom.coe_inj + fun _ _ f => hom_ext <| RingHom.coe_inj ((Types.Small.limitCone (F ⋙ forget _)).π.naturality f) } } createsLimitOfReflectsIso fun _ t => { liftedCone := c @@ -477,10 +463,10 @@ instance : makesLimit := IsLimit.ofFaithful (forget₂ _ RingCat.{u}) (RingCat.limitConeIsLimit.{v, u} (F ⋙ forget₂ CommRingCat.{u} RingCat.{u})) - (fun s : Cone F => ofHom <| + (fun s : Cone F => CommRingCat.ofHom <| (RingCat.limitConeIsLimit.{v, u} (F ⋙ forget₂ CommRingCat.{u} RingCat.{u})).lift - ((forget₂ _ RingCat.{u}).mapCone s)) fun _ => rfl } + ((forget₂ _ RingCat.{u}).mapCone s) |>.hom) fun _ => rfl } /-- A choice of limit cone for a functor into `CommRingCat`. (Generally, you'll just want to use `limit F`.) diff --git a/Mathlib/Algebra/Category/Ring/LinearAlgebra.lean b/Mathlib/Algebra/Category/Ring/LinearAlgebra.lean index 6c0ff2192dde4..d19320fcdcd54 100644 --- a/Mathlib/Algebra/Category/Ring/LinearAlgebra.lean +++ b/Mathlib/Algebra/Category/Ring/LinearAlgebra.lean @@ -29,7 +29,7 @@ lemma nontrivial_of_isPushout_of_isField {A B C D : CommRingCat.{u}} [Nontrivial B] [Nontrivial C] (h : IsPushout f g inl inr) : Nontrivial D := by letI : Field A := hA.toField - algebraize [RingHomClass.toRingHom f, RingHomClass.toRingHom g] + algebraize [f.hom, g.hom] let e : D ≅ .of (B ⊗[A] C) := IsColimit.coconePointUniqueUpToIso h.isColimit (CommRingCat.pushoutCoconeIsColimit A B C) let e' : D ≃ B ⊗[A] C := e.commRingCatIsoToRingEquiv.toEquiv diff --git a/Mathlib/Algebra/Category/Ring/Under/Basic.lean b/Mathlib/Algebra/Category/Ring/Under/Basic.lean index a5e50b4d3c16b..f10cb8dcc9368 100644 --- a/Mathlib/Algebra/Category/Ring/Under/Basic.lean +++ b/Mathlib/Algebra/Category/Ring/Under/Basic.lean @@ -28,11 +28,11 @@ namespace CommRingCat instance : CoeSort (Under R) (Type u) where coe A := A.right -instance (A : Under R) : Algebra R A := RingHom.toAlgebra A.hom +instance (A : Under R) : Algebra R A := RingHom.toAlgebra A.hom.hom /-- Turn a morphism in `Under R` into an algebra homomorphism. -/ def toAlgHom {A B : Under R} (f : A ⟶ B) : A →ₐ[R] B where - __ := f.right + __ := f.right.hom commutes' a := by have : (A.hom ≫ f.right) a = B.hom a := by simp simpa only [Functor.const_obj_obj, Functor.id_obj, CommRingCat.comp_apply] using this @@ -69,7 +69,7 @@ namespace AlgHom /-- Make a morphism in `Under R` from an algebra map. -/ def toUnder {A B : Type u} [CommRing A] [CommRing B] [Algebra R A] [Algebra R B] (f : A →ₐ[R] B) : CommRingCat.mkUnder R A ⟶ CommRingCat.mkUnder R B := - Under.homMk f.toRingHom <| by + Under.homMk (CommRingCat.ofHom f.toRingHom) <| by ext a exact f.commutes' a @@ -135,7 +135,7 @@ def tensorProd : Under R ⥤ Under S where variable (S) in /-- The natural isomorphism `S ⊗[R] A ≅ pushout A.hom (algebraMap R S)` in `Under S`. -/ def tensorProdObjIsoPushoutObj (A : Under R) : - mkUnder S (S ⊗[R] A) ≅ (Under.pushout (algebraMap R S)).obj A := + mkUnder S (S ⊗[R] A) ≅ (Under.pushout (ofHom <| algebraMap R S)).obj A := Under.isoMk (CommRingCat.isPushout_tensorProduct R S A).flip.isoPushout <| by simp only [Functor.const_obj_obj, Under.pushout_obj, Functor.id_obj, Under.mk_right, mkUnder_hom, AlgHom.toRingHom_eq_coe, IsPushout.inr_isoPushout_hom, Under.mk_hom] @@ -143,20 +143,20 @@ def tensorProdObjIsoPushoutObj (A : Under R) : @[reassoc (attr := simp)] lemma pushout_inl_tensorProdObjIsoPushoutObj_inv_right (A : Under R) : - pushout.inl A.hom (algebraMap R S) ≫ (tensorProdObjIsoPushoutObj S A).inv.right = - (CommRingCat.ofHom <| Algebra.TensorProduct.includeRight.toRingHom) := by + pushout.inl A.hom (ofHom <| algebraMap R S) ≫ (tensorProdObjIsoPushoutObj S A).inv.right = + (ofHom <| Algebra.TensorProduct.includeRight.toRingHom) := by simp [tensorProdObjIsoPushoutObj] @[reassoc (attr := simp)] lemma pushout_inr_tensorProdObjIsoPushoutObj_inv_right (A : Under R) : - pushout.inr A.hom (algebraMap R S) ≫ + pushout.inr A.hom (ofHom <| algebraMap R S) ≫ (tensorProdObjIsoPushoutObj S A).inv.right = (CommRingCat.ofHom <| Algebra.TensorProduct.includeLeftRingHom) := by simp [tensorProdObjIsoPushoutObj] variable (R S) in /-- `A ↦ S ⊗[R] A` is naturally isomorphic to `A ↦ pushout A.hom (algebraMap R S)`. -/ -def tensorProdIsoPushout : tensorProd R S ≅ Under.pushout (algebraMap R S) := +def tensorProdIsoPushout : tensorProd R S ≅ Under.pushout (ofHom <| algebraMap R S) := NatIso.ofComponents (fun A ↦ tensorProdObjIsoPushoutObj S A) <| by intro A B f dsimp diff --git a/Mathlib/Algebra/Category/Ring/Under/Limits.lean b/Mathlib/Algebra/Category/Ring/Under/Limits.lean index 24c129198c8b4..c0da3ca2e03cf 100644 --- a/Mathlib/Algebra/Category/Ring/Under/Limits.lean +++ b/Mathlib/Algebra/Category/Ring/Under/Limits.lean @@ -38,8 +38,8 @@ variable {ι : Type u} (P : ι → Under R) /-- The canonical fan on `P : ι → Under R` given by `∀ i, P i`. -/ def piFan : Fan P := - Fan.mk (Under.mk <| ofHom <| Pi.ringHom (fun i ↦ (P i).hom)) - (fun i ↦ Under.homMk (Pi.evalRingHom _ i)) + Fan.mk (Under.mk <| ofHom <| Pi.ringHom (fun i ↦ (P i).hom.hom)) + (fun i ↦ Under.homMk (ofHom <| Pi.evalRingHom _ i)) /-- The canonical fan is limiting. -/ def piFanIsLimit : IsLimit (piFan P) := @@ -204,14 +204,14 @@ variable (f : R ⟶ S) /-- `Under.pushout f` preserves finite products. -/ instance : PreservesFiniteProducts (Under.pushout f) where preserves _ := - letI : Algebra R S := RingHom.toAlgebra f + letI : Algebra R S := f.hom.toAlgebra preservesLimitsOfShape_of_natIso (tensorProdIsoPushout R S) /-- `Under.pushout f` preserves finite limits if `f` is flat. -/ -lemma preservesFiniteLimits_of_flat (hf : RingHom.Flat f) : +lemma preservesFiniteLimits_of_flat (hf : RingHom.Flat f.hom) : PreservesFiniteLimits (Under.pushout f) where preservesFiniteLimits _ := - letI : Algebra R S := RingHom.toAlgebra f + letI : Algebra R S := f.hom.toAlgebra haveI : Module.Flat R S := hf preservesLimitsOfShape_of_natIso (tensorProdIsoPushout R S) diff --git a/Mathlib/Algebra/Ring/Hom/Defs.lean b/Mathlib/Algebra/Ring/Hom/Defs.lean index c03030622109a..1f3bf480a9abf 100644 --- a/Mathlib/Algebra/Ring/Hom/Defs.lean +++ b/Mathlib/Algebra/Ring/Hom/Defs.lean @@ -542,6 +542,9 @@ def id (α : Type*) [NonAssocSemiring α] : α →+* α where instance : Inhabited (α →+* α) := ⟨id α⟩ +@[simp] +theorem coe_id : ⇑(RingHom.id α) = _root_.id := rfl + @[simp] theorem id_apply (x : α) : RingHom.id α x = x := rfl diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index 96002facc9d47..a3aef88d53701 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -313,7 +313,7 @@ lemma isoSpec_hom_base_apply (x : U) : (Iso.eq_comp_inv _).mpr (Scheme.Opens.germ_stalkIso_hom U (V := ⊤) x trivial), X.presheaf.germ_res_assoc, Spec.map_comp, Scheme.comp_base_apply] congr 1 - exact IsLocalRing.comap_closedPoint (U.stalkIso x).inv + exact IsLocalRing.comap_closedPoint (U.stalkIso x).inv.hom lemma isoSpec_inv_appTop : hU.isoSpec.inv.appTop = U.topIso.hom ≫ (Scheme.ΓSpecIso Γ(X, U)).inv := by @@ -549,11 +549,11 @@ theorem exists_basicOpen_le {V : X.Opens} (x : V) (h : ↑x ∈ U) : noncomputable instance {R : CommRingCat} {U} : Algebra R Γ(Spec R, U) := - ((Scheme.ΓSpecIso R).inv ≫ (Spec R).presheaf.map (homOfLE le_top).op).toAlgebra + ((Scheme.ΓSpecIso R).inv ≫ (Spec R).presheaf.map (homOfLE le_top).op).hom.toAlgebra @[simp] lemma algebraMap_Spec_obj {R : CommRingCat} {U} : algebraMap R Γ(Spec R, U) = - (Scheme.ΓSpecIso R).inv ≫ (Spec R).presheaf.map (homOfLE le_top).op := rfl + ((Scheme.ΓSpecIso R).inv ≫ (Spec R).presheaf.map (homOfLE le_top).op).hom := rfl instance {R : CommRingCat} {f : R} : IsLocalization.Away f Γ(Spec R, PrimeSpectrum.basicOpen f) := @@ -586,13 +586,13 @@ theorem isLocalization_basicOpen : -- Porting note: more hand holding is required here, the next 4 lines were not necessary delta StructureSheaf.openAlgebra congr 1 - rw [CommRingCat.ringHom_comp_eq_comp, Iso.commRingIsoToRingEquiv_toRingHom, asIso_hom] + rw [RingEquiv.toRingHom_eq_coe, CategoryTheory.Iso.commRingCatIsoToRingEquiv_toRingHom, asIso_hom] dsimp [CommRingCat.ofHom, RingHom.algebraMap_toAlgebra] - change X.presheaf.map _ ≫ basicOpenSectionsToAffine hU f = _ + change (X.presheaf.map _ ≫ basicOpenSectionsToAffine hU f).hom = _ delta basicOpenSectionsToAffine rw [hU.fromSpec.naturality_assoc, hU.fromSpec_app_self] simp only [Category.assoc, ← Functor.map_comp, ← op_comp] - apply StructureSheaf.toOpen_res + exact CommRingCat.hom_ext_iff.mp (StructureSheaf.toOpen_res _ _ _ _) instance _root_.AlgebraicGeometry.isLocalization_away_of_isAffine [IsAffine X] (r : Γ(X, ⊤)) : @@ -604,29 +604,31 @@ lemma appLE_eq_away_map {X Y : Scheme.{u}} (f : X ⟶ Y) {U : Y.Opens} (hU : IsA letI := hU.isLocalization_basicOpen r letI := hV.isLocalization_basicOpen (f.appLE U V e r) f.appLE (Y.basicOpen r) (X.basicOpen (f.appLE U V e r)) (by simp [Scheme.Hom.appLE]) = - IsLocalization.Away.map _ _ (f.appLE U V e) r := by + CommRingCat.ofHom (IsLocalization.Away.map _ _ (f.appLE U V e).hom r) := by letI := hU.isLocalization_basicOpen r letI := hV.isLocalization_basicOpen (f.appLE U V e r) + ext : 1 apply IsLocalization.ringHom_ext (.powers r) - rw [← CommRingCat.comp_eq_ring_hom_comp, IsLocalization.Away.map, IsLocalization.map_comp, - RingHom.algebraMap_toAlgebra, RingHom.algebraMap_toAlgebra, ← CommRingCat.comp_eq_ring_hom_comp, - Scheme.Hom.appLE_map, Scheme.Hom.map_appLE] + rw [IsLocalization.Away.map, IsLocalization.map_comp, + RingHom.algebraMap_toAlgebra, RingHom.algebraMap_toAlgebra, ← CommRingCat.hom_comp, + ← CommRingCat.hom_comp, Scheme.Hom.appLE_map, Scheme.Hom.map_appLE] lemma app_basicOpen_eq_away_map {X Y : Scheme.{u}} (f : X ⟶ Y) {U : Y.Opens} (hU : IsAffineOpen U) (h : IsAffineOpen (f ⁻¹ᵁ U)) (r : Γ(Y, U)) : haveI := hU.isLocalization_basicOpen r haveI := h.isLocalization_basicOpen (f.app U r) f.app (Y.basicOpen r) = - IsLocalization.Away.map Γ(Y, Y.basicOpen r) Γ(X, X.basicOpen (f.app U r)) (f.app U) r ≫ - X.presheaf.map (eqToHom (by simp)).op := by + (CommRingCat.ofHom + (IsLocalization.Away.map Γ(Y, Y.basicOpen r) Γ(X, X.basicOpen (f.app U r)) (f.app U).hom r) + ≫ X.presheaf.map (eqToHom (by simp)).op) := by haveI := hU.isLocalization_basicOpen r haveI := h.isLocalization_basicOpen (f.app U r) + ext : 1 apply IsLocalization.ringHom_ext (.powers r) - rw [← CommRingCat.comp_eq_ring_hom_comp, ← CommRingCat.comp_eq_ring_hom_comp, - IsLocalization.Away.map, ← Category.assoc] - nth_rw 3 [CommRingCat.comp_eq_ring_hom_comp] - rw [IsLocalization.map_comp, RingHom.algebraMap_toAlgebra, ← CommRingCat.comp_eq_ring_hom_comp, - RingHom.algebraMap_toAlgebra, Category.assoc, ← X.presheaf.map_comp] + rw [IsLocalization.Away.map, CommRingCat.hom_comp, RingHom.comp_assoc, + IsLocalization.map_comp, RingHom.algebraMap_toAlgebra, + RingHom.algebraMap_toAlgebra, ← RingHom.comp_assoc, ← CommRingCat.hom_comp, + ← CommRingCat.hom_comp, ← X.presheaf.map_comp] simp /-- `f.app (Y.basicOpen r)` is isomorphic to map induced on localizations @@ -636,14 +638,14 @@ def appBasicOpenIsoAwayMap {X Y : Scheme.{u}} (f : X ⟶ Y) {U : Y.Opens} haveI := hU.isLocalization_basicOpen r haveI := h.isLocalization_basicOpen (f.app U r) Arrow.mk (f.app (Y.basicOpen r)) ≅ - Arrow.mk (IsLocalization.Away.map Γ(Y, Y.basicOpen r) - Γ(X, X.basicOpen (f.app U r)) (f.app U) r) := + Arrow.mk (CommRingCat.ofHom (IsLocalization.Away.map Γ(Y, Y.basicOpen r) + Γ(X, X.basicOpen (f.app U r)) (f.app U).hom r)) := Arrow.isoMk (Iso.refl _) (X.presheaf.mapIso (eqToIso (by simp)).op) <| by simp [hU.app_basicOpen_eq_away_map f h] include hU in theorem isLocalization_of_eq_basicOpen {V : X.Opens} (i : V ⟶ U) (e : V = X.basicOpen f) : - @IsLocalization.Away _ _ f Γ(X, V) _ (X.presheaf.map i.op).toAlgebra := by + @IsLocalization.Away _ _ f Γ(X, V) _ (X.presheaf.map i.op).hom.toAlgebra := by subst e; convert isLocalization_basicOpen hU f using 3 instance _root_.AlgebraicGeometry.Γ_restrict_isLocalization @@ -711,9 +713,9 @@ theorem isLocalization_stalk' (y : PrimeSpectrum Γ(X, U)) (hy : hU.fromSpec.bas delta IsLocalization.AtPrime StructureSheaf.stalkAlgebra rw [iff_iff_eq] congr 2 - rw [RingHom.algebraMap_toAlgebra] - refine (Scheme.stalkMap_germ hU.fromSpec _ _ hy).trans ?_ - rw [IsAffineOpen.fromSpec_app_self, Category.assoc, TopCat.Presheaf.germ_res] + rw [RingHom.algebraMap_toAlgebra, RingEquiv.toRingHom_eq_coe, + CategoryTheory.Iso.commRingCatIsoToRingEquiv_toRingHom, asIso_hom, ← CommRingCat.hom_comp, + Scheme.stalkMap_germ, IsAffineOpen.fromSpec_app_self, Category.assoc, TopCat.Presheaf.germ_res] rfl -- Porting note: I have split this into two lemmas @@ -791,7 +793,7 @@ localization of `f` away from `r`. -/ noncomputable def SpecMapRestrictBasicOpenIso {R S : CommRingCat} (f : R ⟶ S) (r : R) : Arrow.mk (Spec.map f ∣_ (PrimeSpectrum.basicOpen r)) ≅ - Arrow.mk (Spec.map <| CommRingCat.ofHom (Localization.awayMap f r)) := by + Arrow.mk (Spec.map <| CommRingCat.ofHom (Localization.awayMap f.hom r)) := by letI e₁ : Localization.Away r ≃ₐ[R] Γ(Spec R, basicOpen r) := IsLocalization.algEquiv (Submonoid.powers r) _ _ letI e₂ : Localization.Away (f r) ≃ₐ[S] Γ(Spec S, basicOpen (f r)) := @@ -804,16 +806,17 @@ def SpecMapRestrictBasicOpenIso {R S : CommRingCat} (f : R ⟶ S) (r : R) : (S := (Localization.Away r)) r rw [← cancel_mono (Spec.map (CommRingCat.ofHom (algebraMap R (Localization.Away r))))] simp only [Arrow.mk_left, Arrow.mk_right, Functor.id_obj, Scheme.isoOfEq_rfl, Iso.refl_trans, - Iso.trans_hom, Functor.mapIso_hom, Iso.op_hom, Iso.symm_hom, RingEquiv.toCommRingCatIso_inv, + Iso.trans_hom, Functor.mapIso_hom, Iso.op_hom, Iso.symm_hom, Scheme.Spec_map, Quiver.Hom.unop_op, Arrow.mk_hom, Category.assoc, ← Spec.map_comp] - show _ ≫ Spec.map ((e₂.toRingHom.comp (Localization.awayMap f r)).comp (algebraMap R _)) - = _ ≫ _ ≫ Spec.map (e₁.toRingHom.comp (algebraMap R _)) + show _ ≫ Spec.map (CommRingCat.ofHom + ((e₂.toRingHom.comp (Localization.awayMap f.hom r)).comp (algebraMap R _))) + = _ ≫ _ ≫ Spec.map (CommRingCat.ofHom (e₁.toRingHom.comp (algebraMap R _))) rw [RingHom.comp_assoc] - conv => enter [1,2,1,2]; tactic => exact IsLocalization.map_comp _ + conv => enter [1,2,1,1,2]; tactic => exact IsLocalization.map_comp _ rw [← RingHom.comp_assoc] - conv => enter [1,2,1,1]; tactic => exact e₂.toAlgHom.comp_algebraMap - conv => enter [2,2,2,1]; tactic => exact e₁.toAlgHom.comp_algebraMap + conv => enter [1,2,1,1,1]; tactic => exact e₂.toAlgHom.comp_algebraMap + conv => enter [2,2,2,1,1]; tactic => exact e₁.toAlgHom.comp_algebraMap show _ ≫ Spec.map (f ≫ (Scheme.ΓSpecIso S).inv ≫ (Spec S).presheaf.map (homOfLE le_top).op) = Spec.map f ∣_ basicOpen r ≫ _ ≫ Spec.map ((Scheme.ΓSpecIso R).inv ≫ @@ -932,7 +935,7 @@ variable {X : Scheme.{u}} {A : CommRingCat} is the scheme-theoretic image of `f`. For this quotient as an object of `CommRingCat` see `specTargetImage` below. -/ def specTargetImageIdeal (f : X ⟶ Spec A) : Ideal A := - (RingHom.ker <| (((ΓSpec.adjunction).homEquiv X (op A)).symm f).unop) + (RingHom.ker <| (((ΓSpec.adjunction).homEquiv X (op A)).symm f).unop.hom) /-- If `X ⟶ Spec A` is a morphism of schemes, then `Spec` of `specTargetImage f` is the scheme-theoretic image of `f` and `f` factors as @@ -944,12 +947,13 @@ def specTargetImage (f : X ⟶ Spec A) : CommRingCat := /-- If `f : X ⟶ Spec A` is a morphism of schemes, then `f` factors via the inclusion of `Spec (specTargetImage f)` into `X`. -/ def specTargetImageFactorization (f : X ⟶ Spec A) : X ⟶ Spec (specTargetImage f) := - (ΓSpec.adjunction).homEquiv X (op <| specTargetImage f) (Opposite.op (RingHom.kerLift _)) + (ΓSpec.adjunction).homEquiv X (op <| specTargetImage f) (Opposite.op + (CommRingCat.ofHom (RingHom.kerLift _))) /-- If `f : X ⟶ Spec A` is a morphism of schemes, the induced morphism on spectra of `specTargetImageRingHom f` is the inclusion of the scheme-theoretic image of `f` into `Spec A`. -/ def specTargetImageRingHom (f : X ⟶ Spec A) : A ⟶ specTargetImage f := - Ideal.Quotient.mk (specTargetImageIdeal f) + CommRingCat.ofHom (Ideal.Quotient.mk (specTargetImageIdeal f)) variable (f : X ⟶ Spec A) @@ -959,17 +963,17 @@ lemma specTargetImageRingHom_surjective : Function.Surjective (specTargetImageRi lemma specTargetImageFactorization_app_injective : Function.Injective <| (specTargetImageFactorization f).appTop := by let φ : A ⟶ Γ(X, ⊤) := (((ΓSpec.adjunction).homEquiv X (op A)).symm f).unop - let φ' : specTargetImage f ⟶ Scheme.Γ.obj (op X) := RingHom.kerLift φ + let φ' : specTargetImage f ⟶ Scheme.Γ.obj (op X) := CommRingCat.ofHom (RingHom.kerLift φ.hom) show Function.Injective <| ((ΓSpec.adjunction.homEquiv X _) φ'.op).appTop rw [ΓSpec_adjunction_homEquiv_eq] - apply (RingHom.kerLift_injective φ).comp + apply (RingHom.kerLift_injective φ.hom).comp exact ((ConcreteCategory.isIso_iff_bijective (Scheme.ΓSpecIso _).hom).mp inferInstance).injective @[reassoc (attr := simp)] lemma specTargetImageFactorization_comp : specTargetImageFactorization f ≫ Spec.map (specTargetImageRingHom f) = f := by let φ : A ⟶ Γ(X, ⊤) := (((ΓSpec.adjunction).homEquiv X (op A)).symm f).unop - let φ' : specTargetImage f ⟶ Scheme.Γ.obj (op X) := RingHom.kerLift φ + let φ' : specTargetImage f ⟶ Scheme.Γ.obj (op X) := CommRingCat.ofHom (RingHom.kerLift φ.hom) apply ((ΓSpec.adjunction).homEquiv X (op A)).symm.injective apply Opposite.unop_injective rw [Adjunction.homEquiv_naturality_left_symm, Adjunction.homEquiv_counit] @@ -999,7 +1003,7 @@ def affineTargetImageInclusion (f : X ⟶ Y) : affineTargetImage f ⟶ Y := lemma affineTargetImageInclusion_app_surjective : Function.Surjective <| (affineTargetImageInclusion f).appTop := by simp only [Scheme.comp_coeBase, Opens.map_comp_obj, Opens.map_top, Scheme.comp_app, - CommRingCat.coe_comp, affineTargetImageInclusion] + CommRingCat.hom_comp, affineTargetImageInclusion, RingHom.coe_comp] apply Function.Surjective.comp · haveI : (toMorphismProperty (fun f ↦ Function.Surjective f)).RespectsIso := by rw [← toMorphismProperty_respectsIso_iff] @@ -1009,7 +1013,7 @@ lemma affineTargetImageInclusion_app_surjective : (arrowIsoΓSpecOfIsAffine (specTargetImageRingHom (f ≫ Y.isoSpec.hom))).symm).mpr <| specTargetImageRingHom_surjective (f ≫ Y.isoSpec.hom) · apply Function.Bijective.surjective - apply ConcreteCategory.bijective_of_isIso + exact ConcreteCategory.bijective_of_isIso (Scheme.Hom.app Y.isoSpec.inv ⊤) /-- The induced morphism from `X` to the scheme-theoretic image of a morphism `f : X ⟶ Y` with affine target. -/ @@ -1032,9 +1036,9 @@ section Stalks /-- Variant of `AlgebraicGeometry.localRingHom_comp_stalkIso` for `Spec.map`. -/ @[elementwise] lemma Scheme.localRingHom_comp_stalkIso {R S : CommRingCat.{u}} (f : R ⟶ S) (p : PrimeSpectrum S) : - (StructureSheaf.stalkIso R (PrimeSpectrum.comap f p)).hom ≫ + (StructureSheaf.stalkIso R (PrimeSpectrum.comap f.hom p)).hom ≫ (CommRingCat.ofHom <| Localization.localRingHom - (PrimeSpectrum.comap f p).asIdeal p.asIdeal f rfl) ≫ + (PrimeSpectrum.comap f.hom p).asIdeal p.asIdeal f.hom rfl) ≫ (StructureSheaf.stalkIso S p).inv = (Spec.map f).stalkMap p := AlgebraicGeometry.localRingHom_comp_stalkIso f p @@ -1042,8 +1046,8 @@ lemma Scheme.localRingHom_comp_stalkIso {R S : CommRingCat.{u}} (f : R ⟶ S) (p a prime of `S` is isomorphic to the localized ring homomorphism. -/ def Scheme.arrowStalkMapSpecIso {R S : CommRingCat.{u}} (f : R ⟶ S) (p : PrimeSpectrum S) : Arrow.mk ((Spec.map f).stalkMap p) ≅ Arrow.mk (CommRingCat.ofHom <| Localization.localRingHom - (PrimeSpectrum.comap f p).asIdeal p.asIdeal f rfl) := Arrow.isoMk - (StructureSheaf.stalkIso R (PrimeSpectrum.comap f p)) + (PrimeSpectrum.comap f.hom p).asIdeal p.asIdeal f.hom rfl) := Arrow.isoMk + (StructureSheaf.stalkIso R (PrimeSpectrum.comap f.hom p)) (StructureSheaf.stalkIso S p) <| by rw [← Scheme.localRingHom_comp_stalkIso] simp diff --git a/Mathlib/AlgebraicGeometry/AffineSpace.lean b/Mathlib/AlgebraicGeometry/AffineSpace.lean index ac5248a0b208b..20441b03528a1 100644 --- a/Mathlib/AlgebraicGeometry/AffineSpace.lean +++ b/Mathlib/AlgebraicGeometry/AffineSpace.lean @@ -46,11 +46,12 @@ scoped [AlgebraicGeometry] notation "𝔸("n"; "S")" => AffineSpace n S variable {n} in lemma of_mvPolynomial_int_ext {R} {f g : ℤ[n] ⟶ R} (h : ∀ i, f (.X i) = g (.X i)) : f = g := by - suffices f.comp (MvPolynomial.mapEquiv _ ULift.ringEquiv.symm).toRingHom = - g.comp (MvPolynomial.mapEquiv _ ULift.ringEquiv.symm).toRingHom by + suffices f.hom.comp (MvPolynomial.mapEquiv _ ULift.ringEquiv.symm).toRingHom = + g.hom.comp (MvPolynomial.mapEquiv _ ULift.ringEquiv.symm).toRingHom by ext x - obtain ⟨x, rfl⟩ := (MvPolynomial.mapEquiv _ ULift.ringEquiv.symm).surjective x - exact DFunLike.congr_fun this x + · obtain ⟨x⟩ := x + simpa [-map_intCast, -eq_intCast] using DFunLike.congr_fun this (C x) + · simpa [-map_intCast, -eq_intCast] using DFunLike.congr_fun this (X x) ext1 · exact RingHom.ext_int _ _ · simpa using h _ @@ -73,7 +74,7 @@ Use `homOverEquiv` instead. def toSpecMvPolyIntEquiv : (X ⟶ Spec ℤ[n]) ≃ (n → Γ(X, ⊤)) where toFun f i := f.appTop ((Scheme.ΓSpecIso ℤ[n]).inv (.X i)) invFun v := X.toSpecΓ ≫ Spec.map - (MvPolynomial.eval₂Hom ((algebraMap ℤ _).comp ULift.ringEquiv.toRingHom) v) + (CommRingCat.ofHom (MvPolynomial.eval₂Hom ((algebraMap ℤ _).comp ULift.ringEquiv.toRingHom) v)) left_inv f := by apply (ΓSpec.adjunction.homEquiv _ _).symm.injective apply Quiver.Hom.unop_inj @@ -93,8 +94,9 @@ def toSpecMvPolyIntEquiv : (X ⟶ Spec ℤ[n]) ≃ (n → Γ(X, ⊤)) where TopologicalSpace.Opens.map_comp_obj, TopologicalSpace.Opens.map_top, Scheme.comp_app, Scheme.toSpecΓ_appTop, Scheme.ΓSpecIso_naturality, CommRingCat.comp_apply, CommRingCat.coe_of] - erw [Iso.hom_inv_id_apply] - rw [coe_eval₂Hom, eval₂_X] + -- TODO: why does `simp` not apply this lemma? + rw [CommRingCat.hom_inv_apply] + simp lemma toSpecMvPolyIntEquiv_comp {X Y : Scheme} (f : X ⟶ Y) (g : Y ⟶ Spec ℤ[n]) (i) : toSpecMvPolyIntEquiv n (f ≫ g) i = f.appTop (toSpecMvPolyIntEquiv n g i) := rfl @@ -171,21 +173,23 @@ Also see `AffineSpace.SpecIso`. @[simps (config := .lemmasOnly) hom inv] def isoOfIsAffine [IsAffine S] : 𝔸(n; S) ≅ Spec (.of (MvPolynomial n Γ(S, ⊤))) where - hom := 𝔸(n; S).toSpecΓ ≫ Spec.map (eval₂Hom ((𝔸(n; S) ↘ S).appTop) (coord S)) - inv := homOfVector (Spec.map C ≫ S.isoSpec.inv) + hom := 𝔸(n; S).toSpecΓ ≫ Spec.map (CommRingCat.ofHom + (eval₂Hom ((𝔸(n; S) ↘ S).appTop).hom (coord S))) + inv := homOfVector (Spec.map (CommRingCat.ofHom C) ≫ S.isoSpec.inv) ((Scheme.ΓSpecIso (.of (MvPolynomial n Γ(S, ⊤)))).inv ∘ MvPolynomial.X) hom_inv_id := by ext1 · simp only [Category.assoc, homOfVector_over, Category.id_comp] - rw [← Spec.map_comp_assoc, CommRingCat.comp_eq_ring_hom_comp, eval₂Hom_comp_C, - ← Scheme.toSpecΓ_naturality_assoc] + rw [← Spec.map_comp_assoc, ← CommRingCat.ofHom_comp, eval₂Hom_comp_C, + CommRingCat.ofHom_hom, ← Scheme.toSpecΓ_naturality_assoc] simp [Scheme.isoSpec] · simp only [Category.assoc, Scheme.comp_app, Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, TopologicalSpace.Opens.map_top, Scheme.toSpecΓ_appTop, Scheme.ΓSpecIso_naturality, CommRingCat.comp_apply, homOfVector_appTop_coord, Function.comp_apply, CommRingCat.coe_of, Scheme.id_app, CommRingCat.id_apply] - erw [Iso.hom_inv_id_apply] + -- TODO: why does `simp` not apply this? + rw [CommRingCat.hom_inv_apply] exact eval₂_X _ _ _ inv_hom_id := by apply ext_of_isAffine @@ -193,24 +197,27 @@ def isoOfIsAffine [IsAffine S] : TopologicalSpace.Opens.map_top, Scheme.comp_app, Scheme.toSpecΓ_appTop, Scheme.ΓSpecIso_naturality, Category.assoc, Scheme.id_app, ← Iso.eq_inv_comp, Category.comp_id] + ext : 1 apply ringHom_ext' - · show _ = CommRingCat.ofHom C ≫ _ - rw [CommRingCat.comp_eq_ring_hom_comp, RingHom.comp_assoc, eval₂Hom_comp_C, - ← CommRingCat.comp_eq_ring_hom_comp, ← cancel_mono (Scheme.ΓSpecIso _).hom] + · show _ = (CommRingCat.ofHom C ≫ _).hom + rw [CommRingCat.hom_comp, RingHom.comp_assoc, eval₂Hom_comp_C, + ← CommRingCat.hom_comp, ← CommRingCat.hom_ext_iff, + ← cancel_mono (Scheme.ΓSpecIso _).hom] rw [← Scheme.comp_appTop, homOfVector_over, Scheme.comp_appTop] - simp only [Category.assoc, Scheme.ΓSpecIso_naturality, ← Scheme.toSpecΓ_appTop] + simp only [Category.assoc, Scheme.ΓSpecIso_naturality, CommRingCat.of_carrier, + ← Scheme.toSpecΓ_appTop] rw [← Scheme.comp_appTop_assoc, Scheme.isoSpec, asIso_inv, IsIso.hom_inv_id] simp - rfl · intro i - erw [CommRingCat.comp_apply, coe_eval₂Hom] + rw [CommRingCat.comp_apply, coe_eval₂Hom] simp only [eval₂_X] exact homOfVector_appTop_coord _ _ _ @[simp] lemma isoOfIsAffine_hom_appTop [IsAffine S] : (isoOfIsAffine n S).hom.appTop = - (Scheme.ΓSpecIso _).hom ≫ eval₂Hom ((𝔸(n; S) ↘ S).appTop) (coord S) := by + (Scheme.ΓSpecIso _).hom ≫ CommRingCat.ofHom + (eval₂Hom ((𝔸(n; S) ↘ S).appTop).hom (coord S)) := by simp [isoOfIsAffine_hom] @[simp] @@ -220,7 +227,7 @@ lemma isoOfIsAffine_inv_appTop_coord [IsAffine S] (i) : @[reassoc (attr := simp)] lemma isoOfIsAffine_inv_over [IsAffine S] : - (isoOfIsAffine n S).inv ≫ 𝔸(n; S) ↘ S = Spec.map C ≫ S.isoSpec.inv := + (isoOfIsAffine n S).inv ≫ 𝔸(n; S) ↘ S = Spec.map (CommRingCat.ofHom C) ≫ S.isoSpec.inv := pullback.lift_fst _ _ _ instance [IsAffine S] : IsAffine 𝔸(n; S) := isAffine_of_isIso (isoOfIsAffine n S).hom @@ -235,13 +242,14 @@ def SpecIso (R : CommRingCat.{max u v}) : @[simp] lemma SpecIso_hom_appTop (R : CommRingCat.{max u v}) : (SpecIso n R).hom.appTop = (Scheme.ΓSpecIso _).hom ≫ - eval₂Hom ((Scheme.ΓSpecIso _).inv ≫ (𝔸(n; Spec R) ↘ Spec R).appTop) (coord (Spec R)) := by - simp only [SpecIso, Iso.trans_hom, Functor.mapIso_hom, Iso.op_hom, RingEquiv.toCommRingCatIso_hom, + CommRingCat.ofHom (eval₂Hom ((Scheme.ΓSpecIso _).inv ≫ + (𝔸(n; Spec R) ↘ Spec R).appTop).hom (coord (Spec R))) := by + simp only [SpecIso, Iso.trans_hom, Functor.mapIso_hom, Iso.op_hom, RingEquiv.toRingHom_eq_coe, Scheme.Spec_map, Quiver.Hom.unop_op, Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, TopologicalSpace.Opens.map_top, Scheme.comp_app, - isoOfIsAffine_hom_appTop] - erw [Scheme.ΓSpecIso_naturality_assoc] + isoOfIsAffine_hom_appTop, Scheme.ΓSpecIso_naturality_assoc] congr 1 + ext : 1 apply ringHom_ext' · ext; simp · simp @@ -249,25 +257,24 @@ lemma SpecIso_hom_appTop (R : CommRingCat.{max u v}) : @[simp] lemma SpecIso_inv_appTop_coord (R : CommRingCat.{max u v}) (i) : (SpecIso n R).inv.appTop (coord _ i) = (Scheme.ΓSpecIso (.of _)).inv (.X i) := by - simp only [SpecIso, Iso.trans_inv, Functor.mapIso_inv, Iso.op_inv, RingEquiv.toCommRingCatIso_inv, + simp only [SpecIso, Iso.trans_inv, Functor.mapIso_inv, Iso.op_inv, mapEquiv_symm, RingEquiv.toRingHom_eq_coe, Scheme.Spec_map, Quiver.Hom.unop_op, Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, TopologicalSpace.Opens.map_top, Scheme.comp_app, CommRingCat.comp_apply] - erw [isoOfIsAffine_inv_appTop_coord, ← CommRingCat.comp_apply] - rw [← Scheme.ΓSpecIso_inv_naturality] - erw [CommRingCat.comp_apply] + rw [isoOfIsAffine_inv_appTop_coord, ← CommRingCat.comp_apply, ← Scheme.ΓSpecIso_inv_naturality, + CommRingCat.comp_apply] congr 1 exact map_X _ _ @[reassoc (attr := simp)] lemma SpecIso_inv_over (R : CommRingCat.{max u v}) : - (SpecIso n R).inv ≫ 𝔸(n; Spec R) ↘ Spec R = Spec.map C := by - simp only [SpecIso, Iso.trans_inv, Functor.mapIso_inv, Iso.op_inv, RingEquiv.toCommRingCatIso_inv, + (SpecIso n R).inv ≫ 𝔸(n; Spec R) ↘ Spec R = Spec.map (CommRingCat.ofHom C) := by + simp only [SpecIso, Iso.trans_inv, Functor.mapIso_inv, Iso.op_inv, mapEquiv_symm, RingEquiv.toRingHom_eq_coe, Scheme.Spec_map, Quiver.Hom.unop_op, Category.assoc, isoOfIsAffine_inv_over, Scheme.isoSpec_Spec_inv, ← Spec.map_comp] congr 1 rw [Iso.inv_comp_eq] - ext + ext : 2 exact map_C _ _ section functorial @@ -299,27 +306,26 @@ lemma map_comp {S S' S'' : Scheme} (f : S ⟶ S') (g : S' ⟶ S'') : lemma map_Spec_map {R S : CommRingCat.{max u v}} (φ : R ⟶ S) : map n (Spec.map φ) = - (SpecIso n S).hom ≫ Spec.map (MvPolynomial.map φ) ≫ (SpecIso n R).inv := by + (SpecIso n S).hom ≫ Spec.map (CommRingCat.ofHom (MvPolynomial.map φ.hom)) ≫ + (SpecIso n R).inv := by rw [← Iso.inv_comp_eq] ext1 · simp only [map_over, Category.assoc, SpecIso_inv_over, SpecIso_inv_over_assoc, - ← Spec.map_comp, CommRingCat.comp_eq_ring_hom_comp] - rw [map_comp_C] + ← Spec.map_comp, ← CommRingCat.ofHom_comp] + rw [map_comp_C, CommRingCat.ofHom_comp, CommRingCat.ofHom_hom] · simp only [Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, TopologicalSpace.Opens.map_top, Scheme.comp_app, CommRingCat.comp_apply] conv_lhs => enter[2]; tactic => exact map_appTop_coord _ _ conv_rhs => enter[2]; tactic => exact SpecIso_inv_appTop_coord _ _ - erw [SpecIso_inv_appTop_coord, ← CommRingCat.comp_apply] - rw [← Scheme.ΓSpecIso_inv_naturality] - erw [CommRingCat.comp_apply, map_X] - rfl + rw [SpecIso_inv_appTop_coord, ← CommRingCat.comp_apply, ← Scheme.ΓSpecIso_inv_naturality, + CommRingCat.comp_apply, map_X] /-- The map between affine spaces over affine bases is isomorphic to the natural map between polynomial rings. -/ def mapSpecMap {R S : CommRingCat.{max u v}} (φ : R ⟶ S) : Arrow.mk (map n (Spec.map φ)) ≅ - Arrow.mk (Spec.map (CommRingCat.ofHom (MvPolynomial.map (σ := n) φ))) := - Arrow.isoMk (SpecIso n S) (SpecIso n R) (by simp [map_Spec_map]; rfl) + Arrow.mk (Spec.map (CommRingCat.ofHom (MvPolynomial.map (σ := n) φ.hom))) := + Arrow.isoMk (SpecIso n S) (SpecIso n R) (by simp [map_Spec_map]) /-- `𝔸(n; S)` is functorial wrt `n`. -/ def reindex {n m : Type v} (i : m → n) (S : Scheme.{max u v}) : 𝔸(n; S) ⟶ 𝔸(m; S) := diff --git a/Mathlib/AlgebraicGeometry/FunctionField.lean b/Mathlib/AlgebraicGeometry/FunctionField.lean index 353df65b03ce8..a814d14268a98 100644 --- a/Mathlib/AlgebraicGeometry/FunctionField.lean +++ b/Mathlib/AlgebraicGeometry/FunctionField.lean @@ -42,12 +42,12 @@ noncomputable abbrev Scheme.germToFunctionField [IrreducibleSpace X] (U : X.Open noncomputable instance [IrreducibleSpace X] (U : X.Opens) [Nonempty U] : Algebra Γ(X, U) X.functionField := - (X.germToFunctionField U).toAlgebra + (X.germToFunctionField U).hom.toAlgebra noncomputable instance [IsIntegral X] : Field X.functionField := by refine .ofIsUnitOrEqZero fun a ↦ ?_ - obtain ⟨U, m, s, rfl⟩ := TopCat.Presheaf.germ_exist _ _ a - rw [or_iff_not_imp_right, ← (X.presheaf.germ _ _ m).map_zero] + obtain ⟨U, m, s, rfl⟩ := TopCat.Presheaf.germ_exist (C := CommRingCat) _ _ a + rw [or_iff_not_imp_right, ← (X.presheaf.germ _ _ m).hom.map_zero] intro ha replace ha := ne_of_apply_ne _ ha have hs : genericPoint X ∈ RingedSpace.basicOpen _ s := by @@ -56,14 +56,14 @@ noncomputable instance [IsIntegral X] : Field X.functionField := by · erw [basicOpen_eq_bot_iff] exact ha · exact (RingedSpace.basicOpen _ _).isOpen - have := (X.presheaf.germ _ _ hs).isUnit_map (RingedSpace.isUnit_res_basicOpen _ s) - rwa [TopCat.Presheaf.germ_res_apply] at this + have := (X.presheaf.germ _ _ hs).hom.isUnit_map (RingedSpace.isUnit_res_basicOpen _ s) + rwa [CommRingCat.germ_res_apply] at this theorem germ_injective_of_isIntegral [IsIntegral X] {U : X.Opens} (x : X) (hx : x ∈ U) : Function.Injective (X.presheaf.germ U x hx) := by rw [injective_iff_map_eq_zero] intro y hy - rw [← (X.presheaf.germ U x hx).map_zero] at hy + rw [← (X.presheaf.germ U x hx).hom.map_zero] at hy obtain ⟨W, hW, iU, iV, e⟩ := X.presheaf.germ_eq _ hx hx _ _ hy cases Subsingleton.elim iU iV haveI : Nonempty W := ⟨⟨_, hW⟩⟩ @@ -87,19 +87,21 @@ theorem genericPoint_eq_of_isOpenImmersion {X Y : Scheme} (f : X ⟶ Y) [H : IsO noncomputable instance stalkFunctionFieldAlgebra [IrreducibleSpace X] (x : X) : Algebra (X.presheaf.stalk x) X.functionField := by + -- TODO: can we write this normally after the refactor finishes? apply RingHom.toAlgebra - exact X.presheaf.stalkSpecializes ((genericPoint_spec X).specializes trivial) + exact (X.presheaf.stalkSpecializes ((genericPoint_spec X).specializes trivial)).hom instance functionField_isScalarTower [IrreducibleSpace X] (U : X.Opens) (x : U) [Nonempty U] : IsScalarTower Γ(X, U) (X.presheaf.stalk x) X.functionField := by apply IsScalarTower.of_algebraMap_eq' simp_rw [RingHom.algebraMap_toAlgebra] - change _ = X.presheaf.germ U x x.2 ≫ _ + change _ = (X.presheaf.germ U x x.2 ≫ _).hom rw [X.presheaf.germ_stalkSpecializes] noncomputable instance (R : CommRingCat.{u}) [IsDomain R] : Algebra R (Spec R).functionField := - RingHom.toAlgebra <| by change CommRingCat.of R ⟶ _; apply StructureSheaf.toStalk + -- TODO: can we write this normally after the refactor finishes? + RingHom.toAlgebra <| by apply CommRingCat.Hom.hom; apply StructureSheaf.toStalk @[simp] theorem genericPoint_eq_bot_of_affine (R : CommRingCat) [IsDomain R] : diff --git a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean index 188473407ba08..d50d424421ef6 100644 --- a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean +++ b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean @@ -63,7 +63,7 @@ variable (X : LocallyRingedSpace.{u}) /-- The canonical map from the underlying set to the prime spectrum of `Γ(X)`. -/ def toΓSpecFun : X → PrimeSpectrum (Γ.obj (op X)) := fun x => - comap (X.presheaf.Γgerm x) (IsLocalRing.closedPoint (X.presheaf.stalk x)) + comap (X.presheaf.Γgerm x).hom (IsLocalRing.closedPoint (X.presheaf.stalk x)) theorem not_mem_prime_iff_unit_in_stalk (r : Γ.obj (op X)) (x : X) : r ∉ (X.toΓSpecFun x).asIdeal ↔ IsUnit (X.presheaf.Γgerm x r) := by @@ -114,17 +114,24 @@ abbrev toToΓSpecMapBasicOpen : /-- `r` is a unit as a section on the basic open defined by `r`. -/ theorem isUnit_res_toΓSpecMapBasicOpen : IsUnit (X.toToΓSpecMapBasicOpen r r) := by convert - (X.presheaf.map <| (eqToHom <| X.toΓSpecMapBasicOpen_eq r).op).isUnit_map + (X.presheaf.map <| (eqToHom <| X.toΓSpecMapBasicOpen_eq r).op).hom.isUnit_map (X.toRingedSpace.isUnit_res_basicOpen r) -- Porting note: `rw [comp_apply]` to `erw [comp_apply]` - erw [← comp_apply, ← Functor.map_comp] + erw [← CommRingCat.comp_apply, ← Functor.map_comp] congr /-- Define the sheaf hom on individual basic opens for the unit. -/ def toΓSpecCApp : (structureSheaf <| Γ.obj <| op X).val.obj (op <| basicOpen r) ⟶ X.presheaf.obj (op <| X.toΓSpecMapBasicOpen r) := - IsLocalization.Away.lift r (isUnit_res_toΓSpecMapBasicOpen _ r) + -- note: the explicit type annotations were not needed before + -- https://github.com/leanprover-community/mathlib4/pull/19757 + CommRingCat.ofHom <| + IsLocalization.Away.lift + (R := Γ.obj (op X)) + (S := (structureSheaf ↑(Γ.obj (op X))).val.obj (op (basicOpen r))) + r + (isUnit_res_toΓSpecMapBasicOpen _ r) /-- Characterization of the sheaf hom on basic opens, direction ← (next lemma) is used at various places, but → is not used in this file. -/ @@ -136,11 +143,13 @@ theorem toΓSpecCApp_iff -- Porting Note: Type class problem got stuck in `IsLocalization.Away.AwayMap.lift_comp` -- created instance manually. This replaces the `pick_goal` tactics have loc_inst := IsLocalization.to_basicOpen (Γ.obj (op X)) r + refine CommRingCat.hom_ext_iff.trans ?_ rw [← @IsLocalization.Away.lift_comp _ _ _ _ _ _ _ r loc_inst _ (X.isUnit_res_toΓSpecMapBasicOpen r)] --pick_goal 5; exact is_localization.to_basic_open _ r constructor · intro h + ext : 1 exact IsLocalization.ringHom_ext (Submonoid.powers r) h apply congr_arg @@ -212,7 +221,7 @@ def toΓSpec : X ⟶ Spec.locallyRingedSpaceObj (Γ.obj (op X)) where refine IsLocalization.map_units S (⟨r, ?_⟩ : p.asIdeal.primeCompl) apply (not_mem_prime_iff_unit_in_stalk _ _ _).mpr rw [← toStalk_stalkMap_toΓSpec] - erw [comp_apply, ← he] + erw [CommRingCat.comp_apply, ← he] rw [RingHom.map_mul] exact ht.mul <| (IsLocalization.map_units (R := Γ.obj (op X)) S s).map _ @@ -269,11 +278,12 @@ def identityToΓSpec : 𝟭 LocallyRingedSpace.{u} ⟶ Γ.rightOp ⋙ Spec.toLoc apply LocallyRingedSpace.comp_ring_hom_ext · ext1 x dsimp - show PrimeSpectrum.comap (f.c.app (op ⊤)) (X.toΓSpecFun x) = Y.toΓSpecFun (f.base x) + show PrimeSpectrum.comap (f.c.app (op ⊤)).hom (X.toΓSpecFun x) = Y.toΓSpecFun (f.base x) dsimp [toΓSpecFun] - rw [← IsLocalRing.comap_closedPoint (f.stalkMap x), ← - PrimeSpectrum.comap_comp_apply, ← PrimeSpectrum.comap_comp_apply] - congr 2 + rw [← IsLocalRing.comap_closedPoint (f.stalkMap x).hom, ← + PrimeSpectrum.comap_comp_apply, ← PrimeSpectrum.comap_comp_apply, + ← CommRingCat.hom_comp, ← CommRingCat.hom_comp] + congr 3 exact (PresheafedSpace.stalkMap_germ f.1 ⊤ x trivial).symm · intro r rw [LocallyRingedSpace.comp_c_app, ← Category.assoc] @@ -442,7 +452,7 @@ lemma Scheme.toSpecΓ_preimage_basicOpen (X : Scheme.{u}) (r : Γ(X, ⊤)) : rw [← basicOpen_eq_of_affine, Scheme.preimage_basicOpen, ← Scheme.Hom.appTop] congr rw [Scheme.toSpecΓ_appTop] - exact Iso.inv_hom_id_apply _ _ + exact Iso.inv_hom_id_apply (C := CommRingCat) _ _ -- Warning: this LHS of this lemma breaks the structure-sheaf abstraction. @[reassoc (attr := simp)] diff --git a/Mathlib/AlgebraicGeometry/Modules/Tilde.lean b/Mathlib/AlgebraicGeometry/Modules/Tilde.lean index 763031d694da3..affc194a86b50 100644 --- a/Mathlib/AlgebraicGeometry/Modules/Tilde.lean +++ b/Mathlib/AlgebraicGeometry/Modules/Tilde.lean @@ -84,8 +84,7 @@ theorem isLocallyFraction_pred {U : Opens (PrimeSpectrum.Top R)} noncomputable instance (U : (Opens (PrimeSpectrum.Top R))ᵒᵖ) (x : U.unop): Module ((Spec.structureSheaf R).val.obj U) (Localizations M x):= Module.compHom (R := (Localization.AtPrime x.1.asIdeal)) _ - (StructureSheaf.openToLocalization R U.unop x x.2 : - (Spec.structureSheaf R).val.obj U →+* Localization.AtPrime x.1.asIdeal) + ((StructureSheaf.openToLocalization R U.unop x x.2).hom) @[simp] lemma sections_smul_localizations_def @@ -161,7 +160,7 @@ noncomputable def tilde : (Spec (CommRingCat.of R)).Modules where map := fun {U V} i ↦ ofHom -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(Y := ...)` -- This suggests `restrictScalars` needs to be redesigned. - (Y := (restrictScalars ((Spec (CommRingCat.of R)).ringCatSheaf.val.map i)).obj + (Y := (restrictScalars ((Spec (CommRingCat.of R)).ringCatSheaf.val.map i).hom).obj (of ((Spec (CommRingCat.of R)).ringCatSheaf.val.obj V) (M.tildeInType.val.obj V))) { toFun := M.tildeInType.val.map i map_smul' := by intros; rfl @@ -176,7 +175,7 @@ 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 + ModuleCat.restrictScalars (StructureSheaf.globalSectionsIso R).hom.hom namespace Tilde diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean b/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean index 10aa7de982be9..32767133cd5ef 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean @@ -139,7 +139,7 @@ lemma isAffineOpen_of_isAffineOpen_basicOpen (U) (s : Set Γ(X, U)) (hs : Ideal.span s = ⊤) (hs₂ : ∀ i ∈ s, IsAffineOpen (X.basicOpen i)) : IsAffineOpen U := by apply isAffine_of_isAffineOpen_basicOpen (U.topIso.inv '' s) - · rw [← Ideal.map_span U.topIso.inv, hs, Ideal.map_top] + · rw [← Ideal.map_span U.topIso.inv.hom, hs, Ideal.map_top] · rintro _ ⟨j, hj, rfl⟩ rw [← (Scheme.Opens.ι _).isAffineOpen_iff_of_isOpenImmersion, Scheme.image_basicOpen] simpa [Scheme.Opens.toScheme_presheaf_obj] using hs₂ j hj @@ -158,7 +158,7 @@ instance : HasAffineProperty @IsAffineHom fun X _ _ _ ↦ IsAffine X where rw [Scheme.preimage_basicOpen] exact (isAffineOpen_top X).basicOpen _ · intro X Y _ f S hS hS' - apply_fun Ideal.map (f.appTop) at hS + apply_fun Ideal.map (f.appTop).hom at hS rw [Ideal.map_span, Ideal.map_top] at hS apply isAffine_of_isAffineOpen_basicOpen _ hS have : ∀ i : S, IsAffineOpen (f⁻¹ᵁ Y.basicOpen i.1) := hS' diff --git a/Mathlib/AlgebraicGeometry/Morphisms/AffineAnd.lean b/Mathlib/AlgebraicGeometry/Morphisms/AffineAnd.lean index 632a366ca7236..54716daaae092 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/AffineAnd.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/AffineAnd.lean @@ -35,11 +35,11 @@ variable (Q : ∀ {R S : Type u} [CommRing R] [CommRing S], (R →+* S) → Prop /-- This is the affine target morphism property where the source is affine and the induced map of rings on global sections satisfies `P`. -/ def affineAnd : AffineTargetMorphismProperty := - fun X _ f ↦ IsAffine X ∧ Q (f.appTop) + fun X _ f ↦ IsAffine X ∧ Q (f.appTop).hom @[simp] lemma affineAnd_apply {X Y : Scheme.{u}} (f : X ⟶ Y) [IsAffine Y] : - affineAnd Q f ↔ IsAffine X ∧ Q (f.appTop) := + affineAnd Q f ↔ IsAffine X ∧ Q (f.appTop).hom := Iff.rfl attribute [local simp] AffineTargetMorphismProperty.toProperty_apply @@ -65,9 +65,12 @@ lemma affineAnd_isLocal (hPi : RingHom.RespectsIso Q) (hQl : RingHom.Localizatio · simp only [Scheme.preimage_basicOpen, Opens.map_top] exact (isAffineOpen_top X).basicOpen _ · dsimp only - rw [morphismRestrict_appTop, hPi.cancel_right_isIso, Scheme.Opens.ι_image_top] + rw [morphismRestrict_appTop, CommRingCat.hom_comp, hPi.cancel_right_isIso] + -- Not sure why the `show` fixes the following `rw` complaining about "motive is incorrect" + show Q (Scheme.Hom.app f ((Y.basicOpen r).ι ''ᵁ ⊤)).hom + rw [Scheme.Opens.ι_image_top] rw [(isAffineOpen_top Y).app_basicOpen_eq_away_map f (isAffineOpen_top X), - hPi.cancel_right_isIso, ← Scheme.Hom.appTop] + CommRingCat.hom_comp, hPi.cancel_right_isIso, ← Scheme.Hom.appTop] dsimp only [Opens.map_top] haveI := (isAffineOpen_top X).isLocalization_basicOpen (f.appTop r) apply hQl @@ -76,19 +79,22 @@ lemma affineAnd_isLocal (hPi : RingHom.RespectsIso Q) (hQl : RingHom.Localizatio dsimp [affineAnd] at hf haveI : IsAffine X := by apply isAffine_of_isAffineOpen_basicOpen (f.appTop '' s) - · apply_fun Ideal.map (f.appTop) at hs + · apply_fun Ideal.map (f.appTop).hom at hs rwa [Ideal.map_span, Ideal.map_top] at hs · rintro - ⟨r, hr, rfl⟩ simp_rw [Scheme.preimage_basicOpen] at hf exact (hf ⟨r, hr⟩).left - refine ⟨inferInstance, hQs.ofIsLocalization' hPi (f.appTop) s hs fun a ↦ ?_⟩ + refine ⟨inferInstance, hQs.ofIsLocalization' hPi (f.appTop).hom s hs fun a ↦ ?_⟩ refine ⟨Γ(Y, Y.basicOpen a.val), Γ(X, X.basicOpen (f.appTop a.val)), inferInstance, inferInstance, inferInstance, inferInstance, inferInstance, ?_, ?_⟩ · exact (isAffineOpen_top X).isLocalization_basicOpen (f.appTop a.val) · obtain ⟨_, hf⟩ := hf a - rw [morphismRestrict_appTop, hPi.cancel_right_isIso, Scheme.Opens.ι_image_top] at hf + rw [morphismRestrict_appTop, CommRingCat.hom_comp, hPi.cancel_right_isIso] at hf + -- Not sure why the `show` fixes the following `rw` complaining about "motive is incorrect" + have hf : Q (Scheme.Hom.app f ((Y.basicOpen a.1).ι ''ᵁ ⊤)).hom := hf + rw [Scheme.Opens.ι_image_top] at hf rw [(isAffineOpen_top Y).app_basicOpen_eq_away_map _ (isAffineOpen_top X)] at hf - rwa [hPi.cancel_right_isIso] at hf + rwa [CommRingCat.hom_comp, hPi.cancel_right_isIso] at hf /-- If `P` is stable under base change, so is `affineAnd P`. -/ lemma affineAnd_isStableUnderBaseChange (hQi : RingHom.RespectsIso Q) @@ -102,12 +108,16 @@ lemma affineAnd_isStableUnderBaseChange (hQi : RingHom.RespectsIso Q) lemma targetAffineLocally_affineAnd_iff (hQi : RingHom.RespectsIso Q) {X Y : Scheme.{u}} (f : X ⟶ Y) : targetAffineLocally (affineAnd Q) f ↔ ∀ U : Y.Opens, IsAffineOpen U → - IsAffineOpen (f ⁻¹ᵁ U) ∧ Q (f.app U) := by - simp only [targetAffineLocally, affineAnd_apply, morphismRestrict_app, hQi.cancel_right_isIso] + IsAffineOpen (f ⁻¹ᵁ U) ∧ Q (f.app U).hom := by + simp only [targetAffineLocally, affineAnd_apply, morphismRestrict_app, CommRingCat.hom_comp, + hQi.cancel_right_isIso] refine ⟨fun hf U hU ↦ ?_, fun h U ↦ ?_⟩ · obtain ⟨hfU, hf⟩ := hf ⟨U, hU⟩ - exact ⟨hfU, by rwa [Scheme.Opens.ι_image_top] at hf⟩ + use hfU + have hf : Q (Scheme.Hom.app f (((⟨U, hU⟩ : Y.affineOpens) : Y.Opens).ι ''ᵁ ⊤)).hom := hf + rwa [Scheme.Opens.ι_image_top] at hf · refine ⟨(h U U.2).1, ?_⟩ + show Q (Scheme.Hom.app f ((U : Y.Opens).ι ''ᵁ ⊤)).hom rw [Scheme.Opens.ι_image_top] exact (h U U.2).2 @@ -115,7 +125,7 @@ lemma targetAffineLocally_affineAnd_iff (hQi : RingHom.RespectsIso Q) lemma targetAffineLocally_affineAnd_iff' (hQi : RingHom.RespectsIso Q) {X Y : Scheme.{u}} (f : X ⟶ Y) : targetAffineLocally (affineAnd Q) f ↔ - IsAffineHom f ∧ ∀ U : Y.Opens, IsAffineOpen U → Q (f.app U) := by + IsAffineHom f ∧ ∀ U : Y.Opens, IsAffineOpen U → Q (f.app U).hom := by rw [targetAffineLocally_affineAnd_iff hQi, isAffineHom_iff] aesop @@ -134,7 +144,7 @@ lemma targetAffineLocally_affineAnd_iff_affineLocally (hQ : RingHom.PropertyIsLo intro U have : IsAffine (f ⁻¹ᵁ U) := hf.isAffine_preimage U U.2 rw [HasRingHomProperty.iff_of_isAffine (P := affineLocally Q), - morphismRestrict_appTop, hQ.respectsIso.cancel_right_isIso] + morphismRestrict_appTop, CommRingCat.hom_comp, hQ.respectsIso.cancel_right_isIso] apply h rw [Scheme.Opens.ι_image_top] exact U.2 @@ -213,7 +223,7 @@ lemma HasAffineProperty.affineAnd_iff (P : MorphismProperty Scheme.{u}) (hQs : RingHom.OfLocalizationSpan Q) : HasAffineProperty P (affineAnd Q) ↔ ∀ {X Y : Scheme.{u}} (f : X ⟶ Y), P f ↔ - (IsAffineHom f ∧ ∀ U : Y.Opens, IsAffineOpen U → Q (f.app U)) := by + (IsAffineHom f ∧ ∀ U : Y.Opens, IsAffineOpen U → Q (f.app U).hom) := by simp_rw [isAffineHom_iff] refine ⟨fun h X Y f ↦ ?_, fun h ↦ ⟨affineAnd_isLocal hQi hQl hQs, ?_⟩⟩ · rw [eq_targetAffineLocally P, targetAffineLocally_affineAnd_iff hQi] diff --git a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean index 0145a3045d5db..d0c8fb35f7b07 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean @@ -75,7 +75,7 @@ instance (priority := 900) {X Y : Scheme} (f : X ⟶ Y) [IsClosedImmersion f] : /-- Isomorphisms are closed immersions. -/ instance {X Y : Scheme} (f : X ⟶ Y) [IsIso f] : IsClosedImmersion f where base_closed := Homeomorph.isClosedEmbedding <| TopCat.homeoOfIso (asIso f.base) - surj_on_stalks := fun _ ↦ (ConcreteCategory.bijective_of_isIso _).2 + surj_on_stalks := fun _ ↦ (ConcreteCategory.bijective_of_isIso (C := CommRingCat) _).2 instance : MorphismProperty.IsMultiplicative @IsClosedImmersion where id_mem _ := inferInstance @@ -106,7 +106,7 @@ theorem spec_of_surjective {R S : CommRingCat} (f : R ⟶ S) (h : Function.Surje apply (MorphismProperty.arrow_mk_iso_iff (RingHom.toMorphismProperty (fun f ↦ Function.Surjective f)) (Scheme.arrowStalkMapSpecIso f x)).mpr - exact RingHom.surjective_localRingHom_of_surjective f h x.asIdeal + exact RingHom.surjective_localRingHom_of_surjective f.hom h x.asIdeal /-- For any ideal `I` in a commutative ring `R`, the quotient map `specObj R ⟶ specObj (R ⧸ I)` is a closed immersion. -/ @@ -207,7 +207,7 @@ lemma stalkMap_injective_of_isOpenMap_of_injective [CompactSpace X] simp [Set.preimage_image_eq _ hfinj₁] have h0 (i : 𝒰.J) : (𝒰.map i).appLE _ (W i) (by simp) (φ g) = 0 := by rw [← Scheme.Hom.appLE_map _ _ (homOfLE <| hwle i).op, ← Scheme.Hom.map_appLE _ le_rfl w.op] - simp only [CommRingCat.coe_comp_of, RingHom.coe_comp, Function.comp_apply] + simp only [CommRingCat.comp_apply] erw [hg] simp only [map_zero] have h1 (i : 𝒰.J) : ∃ n, (res i) (φ (s ^ n * g)) = 0 := by @@ -304,7 +304,7 @@ instance (priority := 900) {X Y : Scheme.{u}} (f : X ⟶ Y) [h : IsClosedImmersi exact this _ U.2 obtain ⟨_, hf⟩ := h.isAffine_surjective_of_isAffine rw [HasRingHomProperty.iff_of_isAffine (P := @LocallyOfFiniteType)] - exact RingHom.FiniteType.of_surjective (Scheme.Hom.app f ⊤) hf + exact RingHom.FiniteType.of_surjective (Scheme.Hom.app f ⊤).hom hf /-- A surjective closed immersion is an isomorphism when the target is reduced. -/ lemma isIso_of_isClosedImmersion_of_surjective {X Y : Scheme.{u}} (f : X ⟶ Y) @@ -321,7 +321,7 @@ lemma isIso_of_isClosedImmersion_of_surjective {X Y : Scheme.{u}} (f : X ⟶ Y) apply IsClosedImmersion.isIso_of_injective_of_isAffine obtain ⟨hX, hf⟩ := HasAffineProperty.iff_of_isAffine.mp ‹IsClosedImmersion f› let φ := f.appTop - suffices RingHom.ker φ ≤ nilradical _ by + suffices RingHom.ker φ.hom ≤ nilradical _ by rwa [nilradical_eq_zero, Submodule.zero_eq_bot, le_bot_iff, ← RingHom.injective_iff_ker_eq_bot] at this refine (PrimeSpectrum.zeroLocus_eq_top_iff _).mp ?_ diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean b/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean index d439c6baa3a3a..7a34d9b050ddd 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean @@ -255,7 +255,7 @@ end Topologically /-- `stalkwise P` holds for a morphism if all stalks satisfy `P`. -/ def stalkwise (P : ∀ {R S : Type u} [CommRing R] [CommRing S], (R →+* S) → Prop) : MorphismProperty Scheme.{u} := - fun _ _ f => ∀ x, P (f.stalkMap x) + fun _ _ f => ∀ x, P (f.stalkMap x).hom section Stalkwise @@ -297,7 +297,7 @@ lemma stalkwise_isLocalAtSource_of_respectsIso (hP : RingHom.RespectsIso P) : letI := stalkwise_respectsIso hP apply IsLocalAtSource.mk' · intro X Y f U hf x - rw [Scheme.stalkMap_comp, hP.cancel_right_isIso] + rw [Scheme.stalkMap_comp, CommRingCat.hom_comp, hP.cancel_right_isIso] exact hf _ · intro X Y f ι U hU hf x have hy : x ∈ iSup U := by rw [hU]; trivial @@ -307,10 +307,10 @@ lemma stalkwise_isLocalAtSource_of_respectsIso (hP : RingHom.RespectsIso P) : lemma stalkwise_Spec_map_iff (hP : RingHom.RespectsIso P) {R S : CommRingCat} (φ : R ⟶ S) : stalkwise P (Spec.map φ) ↔ ∀ (p : Ideal S) (_ : p.IsPrime), - P (Localization.localRingHom _ p φ rfl) := by + P (Localization.localRingHom _ p φ.hom rfl) := by have hP' : (RingHom.toMorphismProperty P).RespectsIso := RingHom.toMorphismProperty_respectsIso_iff.mp hP - trans ∀ (p : PrimeSpectrum S), P (Localization.localRingHom _ p.asIdeal φ rfl) + trans ∀ (p : PrimeSpectrum S), P (Localization.localRingHom _ p.asIdeal φ.hom rfl) · exact forall_congr' fun p ↦ (RingHom.toMorphismProperty P).arrow_mk_iso_iff (Scheme.arrowStalkMapSpecIso _ _) · exact ⟨fun H p hp ↦ H ⟨p, hp⟩, fun H p ↦ H p.1 p.2⟩ diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Finite.lean b/Mathlib/AlgebraicGeometry/Morphisms/Finite.lean index 60115fff857f2..dcef9ad69d7b4 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Finite.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Finite.lean @@ -33,11 +33,12 @@ the preimage of any affine open subset of `Y` is affine and the induced ring hom is finite. -/ @[mk_iff] class IsFinite {X Y : Scheme} (f : X ⟶ Y) extends IsAffineHom f : Prop where - finite_app (U : Y.Opens) (hU : IsAffineOpen U) : (f.app U).Finite + finite_app (U : Y.Opens) (hU : IsAffineOpen U) : (f.app U).hom.Finite namespace IsFinite -instance : HasAffineProperty @IsFinite (fun X _ f _ ↦ IsAffine X ∧ RingHom.Finite (f.appTop)) := by +instance : HasAffineProperty @IsFinite + (fun X _ f _ ↦ IsAffine X ∧ RingHom.Finite (f.appTop).hom) := by show HasAffineProperty @IsFinite (affineAnd RingHom.Finite) rw [HasAffineProperty.affineAnd_iff _ RingHom.finite_respectsIso RingHom.finite_localizationPreserves RingHom.finite_ofLocalizationSpan] diff --git a/Mathlib/AlgebraicGeometry/Morphisms/FinitePresentation.lean b/Mathlib/AlgebraicGeometry/Morphisms/FinitePresentation.lean index fd89145c9ec5b..4c3998ab096ac 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/FinitePresentation.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/FinitePresentation.lean @@ -38,7 +38,7 @@ and `V ⊆ f ⁻¹' U`, The induced map `Γ(Y, U) ⟶ Γ(X, V)` is of finite pre class LocallyOfFinitePresentation : Prop where finitePresentation_of_affine_subset : ∀ (U : Y.affineOpens) (V : X.affineOpens) (e : V.1 ≤ f ⁻¹ᵁ U.1), - (f.appLE U V e).FinitePresentation + (f.appLE U V e).hom.FinitePresentation instance : HasRingHomProperty @LocallyOfFinitePresentation RingHom.FinitePresentation where isLocal_ringHomProperty := RingHom.finitePresentation_isLocal diff --git a/Mathlib/AlgebraicGeometry/Morphisms/FiniteType.lean b/Mathlib/AlgebraicGeometry/Morphisms/FiniteType.lean index c5324899336fd..23f3fa844445d 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/FiniteType.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/FiniteType.lean @@ -35,7 +35,7 @@ variable {X Y : Scheme.{u}} (f : X ⟶ Y) @[mk_iff] class LocallyOfFiniteType (f : X ⟶ Y) : Prop where finiteType_of_affine_subset : - ∀ (U : Y.affineOpens) (V : X.affineOpens) (e : V.1 ≤ f ⁻¹ᵁ U.1), (f.appLE U V e).FiniteType + ∀ (U : Y.affineOpens) (V : X.affineOpens) (e : V.1 ≤ f ⁻¹ᵁ U.1), (f.appLE U V e).hom.FiniteType instance : HasRingHomProperty @LocallyOfFiniteType RingHom.FiniteType where isLocal_ringHomProperty := RingHom.finiteType_is_local diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Integral.lean b/Mathlib/AlgebraicGeometry/Morphisms/Integral.lean index 09ee678d31470..c6a2c5393bdfa 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Integral.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Integral.lean @@ -34,12 +34,12 @@ the preimage of any affine open subset of `Y` is affine and the induced ring hom is finite. -/ @[mk_iff] class IsIntegralHom {X Y : Scheme} (f : X ⟶ Y) extends IsAffineHom f : Prop where - integral_app (U : Y.Opens) (hU : IsAffineOpen U) : (f.app U).IsIntegral + integral_app (U : Y.Opens) (hU : IsAffineOpen U) : (f.app U).hom.IsIntegral namespace IsIntegralHom instance hasAffineProperty : HasAffineProperty @IsIntegralHom - fun X _ f _ ↦ IsAffine X ∧ RingHom.IsIntegral (f.app ⊤) := by + fun X _ f _ ↦ IsAffine X ∧ RingHom.IsIntegral (f.app ⊤).hom := by show HasAffineProperty @IsIntegralHom (affineAnd RingHom.IsIntegral) rw [HasAffineProperty.affineAnd_iff _ RingHom.isIntegral_respectsIso RingHom.isIntegral_isStableUnderBaseChange.localizationPreserves diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean index a2505954c2c6e..c1e4acd9220c2 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean @@ -59,7 +59,7 @@ instance : IsLocalAtTarget @IsPreimmersion := instance (priority := 900) {X Y : Scheme} (f : X ⟶ Y) [IsOpenImmersion f] : IsPreimmersion f where base_embedding := f.isOpenEmbedding.isEmbedding - surj_on_stalks _ := (ConcreteCategory.bijective_of_isIso _).2 + surj_on_stalks _ := (ConcreteCategory.bijective_of_isIso (C := CommRingCat) _).2 instance : MorphismProperty.IsMultiplicative @IsPreimmersion where id_mem _ := inferInstance @@ -91,7 +91,8 @@ theorem comp_iff {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [IsPreimmersion g] ⟨fun _ ↦ of_comp f g, fun _ ↦ inferInstance⟩ lemma Spec_map_iff {R S : CommRingCat.{u}} (f : R ⟶ S) : - IsPreimmersion (Spec.map f) ↔ IsEmbedding (PrimeSpectrum.comap f) ∧ f.SurjectiveOnStalks := by + IsPreimmersion (Spec.map f) ↔ IsEmbedding (PrimeSpectrum.comap f.hom) ∧ + f.hom.SurjectiveOnStalks := by haveI : (RingHom.toMorphismProperty <| fun f ↦ Function.Surjective f).RespectsIso := by rw [← RingHom.toMorphismProperty_respectsIso_iff] exact RingHom.surjective_respectsIso @@ -99,7 +100,7 @@ lemma Spec_map_iff {R S : CommRingCat.{u}} (f : R ⟶ S) : rfl lemma mk_Spec_map {R S : CommRingCat.{u}} {f : R ⟶ S} - (h₁ : IsEmbedding (PrimeSpectrum.comap f)) (h₂ : f.SurjectiveOnStalks) : + (h₁ : IsEmbedding (PrimeSpectrum.comap f.hom)) (h₂ : f.hom.SurjectiveOnStalks) : IsPreimmersion (Spec.map f) := (Spec_map_iff f).mpr ⟨h₁, h₂⟩ diff --git a/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean b/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean index a44c42830da74..ae913e2eda078 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean @@ -222,7 +222,7 @@ lemma isClosedMap_iff_specializingMap (f : X ⟶ Y) [QuasiCompact f] : exact this H ⟨_, rfl⟩ obtain ⟨R, rfl⟩ := hX obtain ⟨φ, rfl⟩ := Spec.homEquiv.symm.surjective f - exact PrimeSpectrum.isClosed_image_of_stableUnderSpecialization φ Z hZ H + exact PrimeSpectrum.isClosed_image_of_stableUnderSpecialization φ.hom Z hZ H @[elab_as_elim] theorem compact_open_induction_on {P : X.Opens → Prop} (S : X.Opens) @@ -245,8 +245,9 @@ theorem compact_open_induction_on {P : X.Opens → Prop} (S : X.Opens) theorem exists_pow_mul_eq_zero_of_res_basicOpen_eq_zero_of_isAffineOpen (X : Scheme) {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 + (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).hom] at H #adaptation_note /-- Prior to nightly-2024-09-29, we could use dot notation here: @@ -262,7 +263,8 @@ theorem exists_pow_mul_eq_zero_of_res_basicOpen_eq_zero_of_isAffineOpen (X : Sch `f ^ n * x = 0` for some `n`. -/ theorem exists_pow_mul_eq_zero_of_res_basicOpen_eq_zero_of_isCompact (X : Scheme.{u}) {U : X.Opens} (hU : IsCompact U.1) (x f : Γ(X, U)) - (H : x |_ X.basicOpen f = 0) : ∃ n : ℕ, f ^ n * x = 0 := by + (H : x |_ᵣ (X.basicOpen f) = 0) : + ∃ n : ℕ, f ^ n * x = 0 := by obtain ⟨s, hs, e⟩ := (isCompactOpen_iff_eq_finset_affine_union U.1).mp ⟨hU, U.2⟩ replace e : U = iSup fun i : s => (i : X.Opens) := by ext1; simpa using e @@ -276,12 +278,11 @@ theorem exists_pow_mul_eq_zero_of_res_basicOpen_eq_zero_of_isCompact (X : Scheme exists_pow_mul_eq_zero_of_res_basicOpen_eq_zero_of_isAffineOpen X i.1.2 (X.presheaf.map (homOfLE (h₁ i)).op x) (X.presheaf.map (homOfLE (h₁ i)).op f) ?_ swap - · delta TopCat.Presheaf.restrictOpen TopCat.Presheaf.restrict at H ⊢ - convert congr_arg (X.presheaf.map (homOfLE _).op) H - -- Note: the below was `simp only [← comp_apply]` - · rw [← comp_apply, ← comp_apply] - · simp only [← Functor.map_comp] - rfl + · show (X.presheaf.map (homOfLE _).op) ((X.presheaf.map (homOfLE _).op).hom x) = 0 + have H : (X.presheaf.map (homOfLE _).op) x = 0 := H + convert congr_arg (X.presheaf.map (homOfLE _).op).hom H + · simp only [← CommRingCat.comp_apply, ← Functor.map_comp] + · rfl · rw [map_zero] · simp only [Scheme.basicOpen_res, inf_le_right] choose n hn using H' @@ -292,6 +293,7 @@ theorem exists_pow_mul_eq_zero_of_res_basicOpen_eq_zero_of_isCompact (X : Scheme subst e apply TopCat.Sheaf.eq_of_locally_eq X.sheaf fun i : s => (i : X.Opens) intro i + show _ = (X.sheaf.val.map _) 0 rw [map_zero] apply this intro i @@ -309,12 +311,12 @@ lemma Scheme.isNilpotent_iff_basicOpen_eq_bot_of_isCompact {X : Scheme.{u}} {U : X.Opens} (hU : IsCompact (U : Set X)) (f : Γ(X, U)) : IsNilpotent f ↔ X.basicOpen f = ⊥ := by refine ⟨X.basicOpen_eq_bot_of_isNilpotent U f, fun hf ↦ ?_⟩ - have h : (1 : Γ(X, U)) |_ X.basicOpen f = (0 : Γ(X, X.basicOpen f)) := by + have h : (1 : Γ(X, U)) |_ᵣ (X.basicOpen f) = 0 := by have e : X.basicOpen f ≤ ⊥ := by rw [hf] - rw [← X.presheaf.restrict_restrict e bot_le] + rw [← CommRingCat.presheaf_restrict_restrict X e bot_le] have : Subsingleton Γ(X, ⊥) := CommRingCat.subsingleton_of_isTerminal X.sheaf.isTerminalOfEmpty - rw [Subsingleton.eq_zero (1 |_ ⊥)] + rw [Subsingleton.eq_zero (1 |_ᵣ ⊥)] show X.presheaf.map _ 0 = 0 rw [map_zero] obtain ⟨n, hn⟩ := exists_pow_mul_eq_zero_of_res_basicOpen_eq_zero_of_isCompact X hU 1 f h diff --git a/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean b/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean index 19388a0c54fbd..5ca1ecc198340 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean @@ -185,11 +185,11 @@ theorem QuasiSeparated.of_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [Qua theorem exists_eq_pow_mul_of_isAffineOpen (X : Scheme) (U : X.Opens) (hU : IsAffineOpen U) (f : Γ(X, U)) (x : Γ(X, X.basicOpen f)) : - ∃ (n : ℕ) (y : Γ(X, U)), y |_ X.basicOpen f = (f |_ X.basicOpen f) ^ n * x := by + ∃ (n : ℕ) (y : Γ(X, U)), y |_ᵣ X.basicOpen f = (f |_ᵣ X.basicOpen f) ^ n * x := by have := (hU.isLocalization_basicOpen f).2 obtain ⟨⟨y, _, n, rfl⟩, d⟩ := this x use n, y - delta TopCat.Presheaf.restrictOpen TopCat.Presheaf.restrict + dsimp only [TopCat.Presheaf.restrictOpenCommRingCat_apply] simpa [mul_comm x] using d.symm theorem exists_eq_pow_mul_of_is_compact_of_quasi_separated_space_aux_aux {X : TopCat} @@ -197,49 +197,50 @@ theorem exists_eq_pow_mul_of_is_compact_of_quasi_separated_space_aux_aux {X : To {y₁ : F.obj (op U₁)} {y₂ : F.obj (op U₂)} {f : F.obj (op <| U₁ ⊔ U₂)} {x : F.obj (op U₃)} (h₄₁ : U₄ ≤ U₁) (h₄₂ : U₄ ≤ U₂) (h₅₁ : U₅ ≤ U₁) (h₅₃ : U₅ ≤ U₃) (h₆₂ : U₆ ≤ U₂) (h₆₃ : U₆ ≤ U₃) (h₇₄ : U₇ ≤ U₄) (h₇₅ : U₇ ≤ U₅) (h₇₆ : U₇ ≤ U₆) - (e₁ : y₁ |_ U₅ = (f |_ U₁ |_ U₅) ^ n₁ * x |_ U₅) - (e₂ : y₂ |_ U₆ = (f |_ U₂ |_ U₆) ^ n₂ * x |_ U₆) : - (((f |_ U₁) ^ n₂ * y₁) |_ U₄) |_ U₇ = (((f |_ U₂) ^ n₁ * y₂) |_ U₄) |_ U₇ := by - apply_fun (fun x : F.obj (op U₅) ↦ x |_ U₇) at e₁ - apply_fun (fun x : F.obj (op U₆) ↦ x |_ U₇) at e₂ - dsimp only [TopCat.Presheaf.restrictOpen, TopCat.Presheaf.restrict] at e₁ e₂ ⊢ - simp only [map_mul, map_pow, ← comp_apply, ← op_comp, ← F.map_comp, homOfLE_comp] at e₁ e₂ ⊢ + (e₁ : y₁ |_ᵣ U₅ = (f |_ᵣ U₁ |_ᵣ U₅) ^ n₁ * x |_ᵣ U₅) + (e₂ : y₂ |_ᵣ U₆ = (f |_ᵣ U₂ |_ᵣ U₆) ^ n₂ * x |_ᵣ U₆) : + (((f |_ᵣ U₁) ^ n₂ * y₁) |_ᵣ U₄) |_ᵣ U₇ = (((f |_ᵣ U₂) ^ n₁ * y₂) |_ᵣ U₄) |_ᵣ U₇ := by + apply_fun (fun x : F.obj (op U₅) ↦ x |_ᵣ U₇) at e₁ + apply_fun (fun x : F.obj (op U₆) ↦ x |_ᵣ U₇) at e₂ + dsimp only [TopCat.Presheaf.restrictOpenCommRingCat_apply] at e₁ e₂ ⊢ + simp only [map_mul, map_pow, ← op_comp, ← F.map_comp, homOfLE_comp, ← CommRingCat.comp_apply] + at e₁ e₂ ⊢ rw [e₁, e₂, mul_left_comm] theorem exists_eq_pow_mul_of_is_compact_of_quasi_separated_space_aux (X : Scheme) (S : X.affineOpens) (U₁ U₂ : X.Opens) {n₁ n₂ : ℕ} {y₁ : Γ(X, U₁)} {y₂ : Γ(X, U₂)} {f : Γ(X, U₁ ⊔ U₂)} {x : Γ(X, X.basicOpen f)} (h₁ : S.1 ≤ U₁) (h₂ : S.1 ≤ U₂) - (e₁ : y₁ |_ X.basicOpen (f |_ U₁) = ((f |_ U₁ |_ X.basicOpen _) ^ n₁) * x |_ X.basicOpen _) - (e₂ : y₂ |_ X.basicOpen (f |_ U₂) = ((f |_ U₂ |_ X.basicOpen _) ^ n₂) * x |_ X.basicOpen _) : + (e₁ : y₁ |_ᵣ X.basicOpen (f |_ᵣ U₁) = + ((f |_ᵣ U₁ |_ᵣ X.basicOpen _) ^ n₁) * x |_ᵣ X.basicOpen _) + (e₂ : y₂ |_ᵣ X.basicOpen (f |_ᵣ U₂) = + ((f |_ᵣ U₂ |_ᵣ X.basicOpen _) ^ n₂) * x |_ᵣ X.basicOpen _) : ∃ n : ℕ, ∀ m, n ≤ m → - ((f |_ U₁) ^ (m + n₂) * y₁) |_ S.1 = ((f |_ U₂) ^ (m + n₁) * y₂) |_ S.1 := by + ((f |_ᵣ U₁) ^ (m + n₂) * y₁) |_ᵣ S.1 = ((f |_ᵣ U₂) ^ (m + n₁) * y₂) |_ᵣ S.1 := by obtain ⟨⟨_, n, rfl⟩, e⟩ := (@IsLocalization.eq_iff_exists _ _ _ _ _ _ - (S.2.isLocalization_basicOpen (f |_ S.1)) - (((f |_ U₁) ^ n₂ * y₁) |_ S.1) - (((f |_ U₂) ^ n₁ * y₂) |_ S.1)).mp <| by + (S.2.isLocalization_basicOpen (f |_ᵣ S.1)) + (((f |_ᵣ U₁) ^ n₂ * y₁) |_ᵣ S.1) + (((f |_ᵣ U₂) ^ n₁ * y₂) |_ᵣ S.1)).mp <| by apply exists_eq_pow_mul_of_is_compact_of_quasi_separated_space_aux_aux (e₁ := e₁) (e₂ := e₂) · show X.basicOpen _ ≤ _ - simp only [TopCat.Presheaf.restrictOpen, TopCat.Presheaf.restrict] - repeat rw [Scheme.basicOpen_res] -- Note: used to be part of the `simp only` + simp only [TopCat.Presheaf.restrictOpenCommRingCat_apply, Scheme.basicOpen_res] exact inf_le_inf h₁ le_rfl · show X.basicOpen _ ≤ _ - simp only [TopCat.Presheaf.restrictOpen, TopCat.Presheaf.restrict] - repeat rw [Scheme.basicOpen_res] -- Note: used to be part of the `simp only` + simp only [TopCat.Presheaf.restrictOpenCommRingCat_apply, Scheme.basicOpen_res] exact inf_le_inf h₂ le_rfl use n intros m hm rw [← tsub_add_cancel_of_le hm] - simp only [TopCat.Presheaf.restrictOpen, TopCat.Presheaf.restrict, - pow_add, map_pow, map_mul, ← comp_apply, mul_assoc, ← Functor.map_comp, ← op_comp, homOfLE_comp, - Subtype.coe_mk] at e ⊢ + simp only [TopCat.Presheaf.restrictOpenCommRingCat_apply, + pow_add, map_pow, map_mul, mul_assoc, ← Functor.map_comp, ← op_comp, homOfLE_comp, + Subtype.coe_mk, ← CommRingCat.comp_apply] at e ⊢ rw [e] theorem exists_eq_pow_mul_of_isCompact_of_isQuasiSeparated (X : Scheme.{u}) (U : X.Opens) (hU : IsCompact U.1) (hU' : IsQuasiSeparated U.1) (f : Γ(X, U)) (x : Γ(X, X.basicOpen f)) : - ∃ (n : ℕ) (y : Γ(X, U)), y |_ X.basicOpen f = (f |_ X.basicOpen f) ^ n * x := by - delta TopCat.Presheaf.restrictOpen TopCat.Presheaf.restrict + ∃ (n : ℕ) (y : Γ(X, U)), y |_ᵣ X.basicOpen f = (f |_ᵣ X.basicOpen f) ^ n * x := by + dsimp only [TopCat.Presheaf.restrictOpenCommRingCat_apply] revert hU' f x refine compact_open_induction_on U hU ?_ ?_ · intro _ f x @@ -260,7 +261,7 @@ theorem exists_eq_pow_mul_of_isCompact_of_isQuasiSeparated (X : Scheme.{u}) (U : obtain ⟨n₂, y₂, hy₂⟩ := exists_eq_pow_mul_of_isAffineOpen X _ U.2 (X.presheaf.map (homOfLE le_sup_right).op f) (X.presheaf.map (homOfLE _).op x) - delta TopCat.Presheaf.restrictOpen TopCat.Presheaf.restrict at hy₂ + dsimp only [TopCat.Presheaf.restrictOpenCommRingCat_apply] at hy₂ -- swap; · rw [X.basicOpen_res]; exact inf_le_right -- Since `S ∪ U` is quasi-separated, `S ∩ U` can be covered by finite affine opens. obtain ⟨s, hs', hs⟩ := @@ -301,9 +302,9 @@ theorem exists_eq_pow_mul_of_isCompact_of_isQuasiSeparated (X : Scheme.{u}) (U : exact @le_iSup X.Opens s _ (fun (i : s) => (i : X.Opens)) i · exact le_of_eq hs · intro i - delta Scheme.sheaf SheafedSpace.sheaf - repeat rw [← comp_apply,] - simp only [← Functor.map_comp, ← op_comp] + -- This unfolds `X.sheaf` and ensures we use `CommRingCat.hom` to apply the morphism + show (X.presheaf.map _) _ = (X.presheaf.map _) _ + simp only [← CommRingCat.comp_apply, ← Functor.map_comp, ← op_comp] apply hn exact Finset.le_sup (Finset.mem_univ _) use Finset.univ.sup n + n₁ + n₂ @@ -312,17 +313,15 @@ theorem exists_eq_pow_mul_of_isCompact_of_isQuasiSeparated (X : Scheme.{u}) (U : use (X.sheaf.objSupIsoProdEqLocus S U.1).inv ⟨⟨_ * _, _ * _⟩, this⟩ refine (X.sheaf.objSupIsoProdEqLocus_inv_eq_iff _ _ _ (X.basicOpen_res _ (homOfLE le_sup_left).op) (X.basicOpen_res _ (homOfLE le_sup_right).op)).mpr ⟨?_, ?_⟩ - · delta Scheme.sheaf SheafedSpace.sheaf + · -- This unfolds `X.sheaf` and ensures we use `CommRingCat.hom` to apply the morphism + show (X.presheaf.map _) _ = (X.presheaf.map _) _ rw [add_assoc, add_comm n₁] - simp only [pow_add, map_pow, map_mul] - rw [hy₁] -- Note: `simp` can't use this - repeat rw [← comp_apply] -- Note: `simp` can't use this - simp only [← mul_assoc, ← Functor.map_comp, ← op_comp, homOfLE_comp] - · delta Scheme.sheaf SheafedSpace.sheaf - simp only [pow_add, map_pow, map_mul] - rw [hy₂] -- Note: `simp` can't use this - repeat rw [← comp_apply] -- Note: `simp` can't use this - simp only [← mul_assoc, ← Functor.map_comp, ← op_comp, homOfLE_comp] + simp only [pow_add, map_pow, map_mul, hy₁, ← CommRingCat.comp_apply, ← mul_assoc, + ← Functor.map_comp, ← op_comp, homOfLE_comp] + · -- This unfolds `X.sheaf` and ensures we use `CommRingCat.hom` to apply the morphism + show (X.presheaf.map _) _ = (X.presheaf.map _) _ + simp only [pow_add, map_pow, map_mul, hy₂, ← CommRingCat.comp_apply, ← mul_assoc, + ← Functor.map_comp, ← op_comp, homOfLE_comp] /-- If `U` is qcqs, then `Γ(X, D(f)) ≃ Γ(X, U)_f` for every `f : Γ(X, U)`. This is known as the **Qcqs lemma** in [R. Vakil, *The rising sea*][RisingSea]. -/ @@ -348,19 +347,19 @@ theorem is_localization_basicOpen_of_qcqs {X : Scheme} {U : X.Opens} (hU : IsCom lemma exists_of_res_eq_of_qcqs {X : Scheme.{u}} {U : TopologicalSpace.Opens X} (hU : IsCompact U.carrier) (hU' : IsQuasiSeparated U.carrier) - {f g s : Γ(X, U)} (hfg : f |_ X.basicOpen s = g |_ X.basicOpen s) : + {f g s : Γ(X, U)} (hfg : f |_ᵣ X.basicOpen s = g |_ᵣ X.basicOpen s) : ∃ n, s ^ n * f = s ^ n * g := by obtain ⟨n, hc⟩ := (is_localization_basicOpen_of_qcqs hU hU' s).exists_of_eq s hfg use n lemma exists_of_res_eq_of_qcqs_of_top {X : Scheme.{u}} [CompactSpace X] [QuasiSeparatedSpace X] - {f g s : Γ(X, ⊤)} (hfg : f |_ X.basicOpen s = g |_ X.basicOpen s) : + {f g s : Γ(X, ⊤)} (hfg : f |_ᵣ X.basicOpen s = g |_ᵣ X.basicOpen s) : ∃ n, s ^ n * f = s ^ n * g := exists_of_res_eq_of_qcqs (U := ⊤) CompactSpace.isCompact_univ isQuasiSeparated_univ hfg lemma exists_of_res_zero_of_qcqs {X : Scheme.{u}} {U : TopologicalSpace.Opens X} (hU : IsCompact U.carrier) (hU' : IsQuasiSeparated U.carrier) - {f s : Γ(X, U)} (hf : f |_ X.basicOpen s = 0) : + {f s : Γ(X, U)} (hf : f |_ᵣ X.basicOpen s = 0) : ∃ n, s ^ n * f = 0 := by suffices h : ∃ n, s ^ n * f = s ^ n * 0 by simpa using h @@ -368,7 +367,7 @@ lemma exists_of_res_zero_of_qcqs {X : Scheme.{u}} {U : TopologicalSpace.Opens X} simpa lemma exists_of_res_zero_of_qcqs_of_top {X : Scheme} [CompactSpace X] [QuasiSeparatedSpace X] - {f s : Γ(X, ⊤)} (hf : f |_ X.basicOpen s = 0) : + {f s : Γ(X, ⊤)} (hf : f |_ᵣ X.basicOpen s = 0) : ∃ n, s ^ n * f = 0 := exists_of_res_zero_of_qcqs (U := ⊤) CompactSpace.isCompact_univ isQuasiSeparated_univ hf @@ -385,7 +384,9 @@ theorem isIso_ΓSpec_adjunction_unit_app_basicOpen {X : Scheme} [CompactSpace X] · refine is_localization_basicOpen_of_qcqs ?_ ?_ _ · exact isCompact_univ · exact isQuasiSeparated_univ - · rw [← CommRingCat.comp_eq_ring_hom_comp] + · simp only [RingHom.algebraMap_toAlgebra] + -- This `rw` doesn't fire as a `simp` (`only`). + rw [← CommRingCat.hom_comp] simp [RingHom.algebraMap_toAlgebra, ← Functor.map_comp] end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean index 952782f698cdd..dd2bb6a76125e 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean @@ -69,18 +69,20 @@ variable (P : ∀ {R S : Type u} [CommRing R] [CommRing S], (R →+* S) → Prop theorem IsStableUnderBaseChange.pullback_fst_appTop (hP : IsStableUnderBaseChange P) (hP' : RespectsIso P) {X Y S : Scheme} [IsAffine X] [IsAffine Y] [IsAffine S] (f : X ⟶ S) (g : Y ⟶ S) - (H : P g.appTop) : P (pullback.fst f g).appTop := by + (H : P g.appTop.hom) : P (pullback.fst f g).appTop.hom := by -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11224): change `rw` to `erw` erw [← PreservesPullback.iso_inv_fst AffineScheme.forgetToScheme (AffineScheme.ofHom f) (AffineScheme.ofHom g)] - rw [Scheme.comp_appTop, hP'.cancel_right_isIso, AffineScheme.forgetToScheme_map] + rw [Scheme.comp_appTop, CommRingCat.hom_comp, hP'.cancel_right_isIso, + AffineScheme.forgetToScheme_map] have := congr_arg Quiver.Hom.unop (PreservesPullback.iso_hom_fst AffineScheme.Γ.rightOp (AffineScheme.ofHom f) (AffineScheme.ofHom g)) simp only [AffineScheme.Γ, Functor.rightOp_obj, Functor.comp_obj, Functor.op_obj, unop_comp, AffineScheme.forgetToScheme_obj, Scheme.Γ_obj, Functor.rightOp_map, Functor.comp_map, Functor.op_map, Quiver.Hom.unop_op, AffineScheme.forgetToScheme_map, Scheme.Γ_map] at this - rw [← this, hP'.cancel_right_isIso, ← pushoutIsoUnopPullback_inl_hom, hP'.cancel_right_isIso] + rw [← this, CommRingCat.hom_comp, hP'.cancel_right_isIso, ← pushoutIsoUnopPullback_inl_hom, + CommRingCat.hom_comp, hP'.cancel_right_isIso] exact hP.pushout_inl _ hP' _ _ H @[deprecated (since := "2024-11-23")] @@ -98,7 +100,7 @@ variable (P : ∀ {R S : Type u} [CommRing R] [CommRing S], (R →+* S) → Prop /-- For `P` a property of ring homomorphisms, `sourceAffineLocally P` holds for `f : X ⟶ Y` whenever `P` holds for the restriction of `f` on every affine open subset of `X`. -/ def sourceAffineLocally : AffineTargetMorphismProperty := fun X _ f _ => - ∀ U : X.affineOpens, P (f.appLE ⊤ U le_top) + ∀ U : X.affineOpens, P (f.appLE ⊤ U le_top).hom /-- For `P` a property of ring homomorphisms, `affineLocally P` holds for `f : X ⟶ Y` if for each affine open `U = Spec A ⊆ Y` and `V = Spec B ⊆ f ⁻¹' U`, the ring hom `A ⟶ B` satisfies `P`. @@ -114,10 +116,10 @@ theorem sourceAffineLocally_respectsIso (h₁ : RingHom.RespectsIso P) : inferInstanceAs (IsIso (e.hom.app _ ≫ X.presheaf.map (eqToHom (e.hom.preimage_image_eq _).symm).op)) rw [← Scheme.appLE_comp_appLE _ _ ⊤ (e.hom ''ᵁ U) U.1 le_top (e.hom.preimage_image_eq _).ge, - h₁.cancel_right_isIso] + CommRingCat.hom_comp, h₁.cancel_right_isIso] exact H ⟨_, U.prop.image_of_isOpenImmersion e.hom⟩ · introv H U - rw [Scheme.comp_appLE, h₁.cancel_left_isIso] + rw [Scheme.comp_appLE, CommRingCat.hom_comp, h₁.cancel_left_isIso] exact H U theorem affineLocally_respectsIso (h : RingHom.RespectsIso P) : (affineLocally P).RespectsIso := @@ -128,17 +130,17 @@ open Scheme in theorem sourceAffineLocally_morphismRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (hU : IsAffineOpen U) : @sourceAffineLocally P _ _ (f ∣_ U) hU ↔ - ∀ (V : X.affineOpens) (e : V.1 ≤ f ⁻¹ᵁ U), P (f.appLE U V e) := by + ∀ (V : X.affineOpens) (e : V.1 ≤ f ⁻¹ᵁ U), P (f.appLE U V e).hom := by dsimp only [sourceAffineLocally] simp only [morphismRestrict_appLE] rw [(affineOpensRestrict (f ⁻¹ᵁ U)).forall_congr_left, Subtype.forall] refine forall₂_congr fun V h ↦ ?_ have := (affineOpensRestrict (f ⁻¹ᵁ U)).apply_symm_apply ⟨V, h⟩ - exact f.appLE_congr _ (Opens.ι_image_top _) congr($(this).1.1) _ + exact f.appLE_congr _ (Opens.ι_image_top _) congr($(this).1.1) (fun f => P f.hom) theorem affineLocally_iff_affineOpens_le {X Y : Scheme.{u}} (f : X ⟶ Y) : affineLocally.{u} P f ↔ - ∀ (U : Y.affineOpens) (V : X.affineOpens) (e : V.1 ≤ f ⁻¹ᵁ U.1), P (f.appLE U V e) := + ∀ (U : Y.affineOpens) (V : X.affineOpens) (e : V.1 ≤ f ⁻¹ᵁ U.1), P (f.appLE U V e).hom := forall_congr' fun U ↦ sourceAffineLocally_morphismRestrict P f U U.2 theorem sourceAffineLocally_isLocal (h₁ : RingHom.RespectsIso P) @@ -150,12 +152,13 @@ theorem sourceAffineLocally_isLocal (h₁ : RingHom.RespectsIso P) rw [sourceAffineLocally_morphismRestrict] intro U hU have : X.basicOpen (f.appLE ⊤ U (by simp) r) = U := by - simp only [Scheme.Hom.appLE, Opens.map_top, CommRingCat.coe_comp_of, RingHom.coe_comp, + simp only [Scheme.Hom.appLE, Opens.map_top, CommRingCat.comp_apply, RingHom.coe_comp, Function.comp_apply] rw [Scheme.basicOpen_res] simpa using hU - rw [← f.appLE_congr _ rfl this P, + rw [← f.appLE_congr _ rfl this (fun f => P f.hom), IsAffineOpen.appLE_eq_away_map f (isAffineOpen_top Y) U.2 _ r] + simp only apply (config := { allowSynthFailures := true }) h₂ exact H U · introv hs hs' U @@ -177,10 +180,11 @@ over some basic open of `U₁` (resp. `V₁`). -/ lemma exists_basicOpen_le_appLE_of_appLE_of_isAffine (hPa : StableUnderCompositionWithLocalizationAwayTarget P) (hPl : LocalizationAwayPreserves P) (x : X) (U₁ : Y.affineOpens) (U₂ : Y.affineOpens) (V₁ : X.affineOpens) (V₂ : X.affineOpens) - (hx₁ : x ∈ V₁.1) (hx₂ : x ∈ V₂.1) (e₂ : V₂.1 ≤ f ⁻¹ᵁ U₂.1) (h₂ : P (f.appLE U₂ V₂ e₂)) + (hx₁ : x ∈ V₁.1) (hx₂ : x ∈ V₂.1) (e₂ : V₂.1 ≤ f ⁻¹ᵁ U₂.1) (h₂ : P (f.appLE U₂ V₂ e₂).hom) (hfx₁ : f.base x ∈ U₁.1) : ∃ (r : Γ(Y, U₁)) (s : Γ(X, V₁)) (_ : x ∈ X.basicOpen s) - (e : X.basicOpen s ≤ f ⁻¹ᵁ Y.basicOpen r), P (f.appLE (Y.basicOpen r) (X.basicOpen s) e) := by + (e : X.basicOpen s ≤ f ⁻¹ᵁ Y.basicOpen r), + P (f.appLE (Y.basicOpen r) (X.basicOpen s) e).hom := by obtain ⟨r, r', hBrr', hBfx⟩ := exists_basicOpen_le_affine_inter U₁.2 U₂.2 (f.base x) ⟨hfx₁, e₂ hx₂⟩ have ha : IsAffineOpen (X.basicOpen (f.appLE U₂ V₂ e₂ r')) := V₂.2.basicOpen _ @@ -196,12 +200,12 @@ lemma exists_basicOpen_le_appLE_of_appLE_of_isAffine simp [Scheme.Hom.appLE] have heq : f.appLE (Y.basicOpen r') (X.basicOpen s') (hBrr' ▸ hBss' ▸ ers) = f.appLE (Y.basicOpen r') (X.basicOpen (f.appLE U₂ V₂ e₂ r')) (by simp [Scheme.Hom.appLE]) ≫ - algebraMap _ _ := by + CommRingCat.ofHom (algebraMap _ _) := by simp only [Scheme.Hom.appLE, homOfLE_leOfHom, CommRingCat.comp_apply, Category.assoc] congr apply X.presheaf.map_comp refine ⟨r, s, hBx, ers, ?_⟩ - · rw [f.appLE_congr _ hBrr' hBss' P, heq] + · rw [f.appLE_congr _ hBrr' hBss' (fun f => P f.hom), heq] apply hPa _ s' _ rw [U₂.2.appLE_eq_away_map f V₂.2] exact hPl _ _ _ _ h₂ @@ -212,10 +216,10 @@ open neighborhoods of `x` (resp. `f.base x`), then `P` also holds for `f` over s lemma exists_affineOpens_le_appLE_of_appLE (hPa : StableUnderCompositionWithLocalizationAwayTarget P) (hPl : LocalizationAwayPreserves P) (x : X) (U₁ : Y.Opens) (U₂ : Y.affineOpens) (V₁ : X.Opens) (V₂ : X.affineOpens) - (hx₁ : x ∈ V₁) (hx₂ : x ∈ V₂.1) (e₂ : V₂.1 ≤ f ⁻¹ᵁ U₂.1) (h₂ : P (f.appLE U₂ V₂ e₂)) + (hx₁ : x ∈ V₁) (hx₂ : x ∈ V₂.1) (e₂ : V₂.1 ≤ f ⁻¹ᵁ U₂.1) (h₂ : P (f.appLE U₂ V₂ e₂).hom) (hfx₁ : f.base x ∈ U₁.1) : ∃ (U' : Y.affineOpens) (V' : X.affineOpens) (_ : U'.1 ≤ U₁) (_ : V'.1 ≤ V₁) (_ : x ∈ V'.1) - (e : V'.1 ≤ f⁻¹ᵁ U'.1), P (f.appLE U' V' e) := by + (e : V'.1 ≤ f⁻¹ᵁ U'.1), P (f.appLE U' V' e).hom := by obtain ⟨r, hBr, hBfx⟩ := U₂.2.exists_basicOpen_le ⟨f.base x, hfx₁⟩ (e₂ hx₂) obtain ⟨s, hBs, hBx⟩ := V₂.2.exists_basicOpen_le ⟨x, hx₁⟩ hx₂ obtain ⟨r', s', hBx', e', hf'⟩ := exists_basicOpen_le_appLE_of_appLE_of_isAffine hPa hPl x @@ -268,11 +272,11 @@ lemma HasAffineProperty : HasAffineProperty P (sourceAffineLocally Q) where /- This is only `inferInstance` because of the `@[local instance]` on `HasAffineProperty` above. -/ instance (priority := 900) : IsLocalAtTarget P := inferInstance -theorem appLE (H : P f) (U : Y.affineOpens) (V : X.affineOpens) (e) : Q (f.appLE U V e) := by +theorem appLE (H : P f) (U : Y.affineOpens) (V : X.affineOpens) (e) : Q (f.appLE U V e).hom := by rw [eq_affineLocally P, affineLocally_iff_affineOpens_le] at H exact H _ _ _ -theorem appTop (H : P f) [IsAffine X] [IsAffine Y] : Q f.appTop := by +theorem appTop (H : P f) [IsAffine X] [IsAffine Y] : Q f.appTop.hom := by rw [Scheme.Hom.appTop, Scheme.Hom.app_eq_appLE] exact appLE P f H ⟨_, isAffineOpen_top _⟩ ⟨_, isAffineOpen_top _⟩ _ @@ -288,16 +292,17 @@ theorem comp_of_isOpenImmersion [IsOpenImmersion f] (H : P g) : X.presheaf.map (eqToHom (f.preimage_image_eq _).symm).op)) rw [← Scheme.appLE_comp_appLE _ _ _ (f ''ᵁ V) V.1 (Set.image_subset_iff.mpr e) (f.preimage_image_eq _).ge, + CommRingCat.hom_comp, (isLocal_ringHomProperty P).respectsIso.cancel_right_isIso] exact H _ ⟨_, V.2.image_of_isOpenImmersion _⟩ _ variable {P f} -lemma iff_appLE : P f ↔ ∀ (U : Y.affineOpens) (V : X.affineOpens) (e), Q (f.appLE U V e) := by +lemma iff_appLE : P f ↔ ∀ (U : Y.affineOpens) (V : X.affineOpens) (e), Q (f.appLE U V e).hom := by rw [eq_affineLocally P, affineLocally_iff_affineOpens_le] theorem of_source_openCover [IsAffine Y] - (𝒰 : X.OpenCover) [∀ i, IsAffine (𝒰.obj i)] (H : ∀ i, Q ((𝒰.map i ≫ f).appTop)) : + (𝒰 : X.OpenCover) [∀ i, IsAffine (𝒰.obj i)] (H : ∀ i, Q ((𝒰.map i ≫ f).appTop.hom)) : P f := by rw [HasAffineProperty.iff_of_isAffine (P := P)] intro U @@ -313,43 +318,43 @@ theorem of_source_openCover [IsAffine Y] (isLocal_ringHomProperty P).respectsIso _ _ hs rintro r refine ⟨_, _, _, IsAffineOpen.isLocalization_basicOpen U.2 r, ?_⟩ - rw [RingHom.algebraMap_toAlgebra, ← CommRingCat.comp_eq_ring_hom_comp, Scheme.Hom.appLE_map] + rw [RingHom.algebraMap_toAlgebra, ← CommRingCat.hom_comp, Scheme.Hom.appLE_map] exact H r | hU i => specialize H i rw [← (isLocal_ringHomProperty P).respectsIso.cancel_right_isIso _ ((IsOpenImmersion.isoOfRangeEq (𝒰.map i) (S i).1.ι - Subtype.range_coe.symm).inv.app _), ← Scheme.comp_appTop, + Subtype.range_coe.symm).inv.app _), ← CommRingCat.hom_comp, ← Scheme.comp_appTop, IsOpenImmersion.isoOfRangeEq_inv_fac_assoc, Scheme.comp_appTop, Scheme.Opens.ι_appTop, Scheme.Hom.appTop, Scheme.Hom.app_eq_appLE, Scheme.Hom.appLE_map] at H - exact (f.appLE_congr _ rfl (by simp) Q).mp H + exact (f.appLE_congr _ rfl (by simp) (fun f => Q f.hom)).mp H theorem iff_of_source_openCover [IsAffine Y] (𝒰 : X.OpenCover) [∀ i, IsAffine (𝒰.obj i)] : - P f ↔ ∀ i, Q ((𝒰.map i ≫ f).appTop) := + P f ↔ ∀ i, Q ((𝒰.map i ≫ f).appTop).hom := ⟨fun H i ↦ appTop P _ (comp_of_isOpenImmersion P (𝒰.map i) f H), of_source_openCover 𝒰⟩ theorem iff_of_isAffine [IsAffine X] [IsAffine Y] : - P f ↔ Q (f.appTop) := by + P f ↔ Q (f.appTop).hom := by rw [iff_of_source_openCover (P := P) (Scheme.coverOfIsIso.{u} (𝟙 _))] simp theorem Spec_iff {R S : CommRingCat.{u}} {φ : R ⟶ S} : - P (Spec.map φ) ↔ Q φ := by + P (Spec.map φ) ↔ Q φ.hom := by have H := (isLocal_ringHomProperty P).respectsIso rw [iff_of_isAffine (P := P), ← H.cancel_right_isIso _ (Scheme.ΓSpecIso _).hom, - Scheme.ΓSpecIso_naturality, H.cancel_left_isIso] + ← CommRingCat.hom_comp, Scheme.ΓSpecIso_naturality, CommRingCat.hom_comp, H.cancel_left_isIso] theorem of_iSup_eq_top [IsAffine Y] {ι : Type*} (U : ι → X.affineOpens) (hU : ⨆ i, (U i : Opens X) = ⊤) - (H : ∀ i, Q (f.appLE ⊤ (U i).1 le_top)) : + (H : ∀ i, Q (f.appLE ⊤ (U i).1 le_top).hom) : P f := by have (i) : IsAffine ((X.openCoverOfISupEqTop _ hU).obj i) := (U i).2 refine of_source_openCover (X.openCoverOfISupEqTop _ hU) fun i ↦ ?_ - simpa [Scheme.Hom.app_eq_appLE] using (f.appLE_congr _ rfl (by simp) Q).mp (H i) + simpa [Scheme.Hom.app_eq_appLE] using (f.appLE_congr _ rfl (by simp) (fun f => Q f.hom)).mp (H i) theorem iff_of_iSup_eq_top [IsAffine Y] {ι : Type*} (U : ι → X.affineOpens) (hU : ⨆ i, (U i : Opens X) = ⊤) : - P f ↔ ∀ i, Q (f.appLE ⊤ (U i).1 le_top) := + P f ↔ ∀ i, Q (f.appLE ⊤ (U i).1 le_top).hom := ⟨fun H _ ↦ appLE P f H ⟨_, isAffineOpen_top _⟩ _ le_top, of_iSup_eq_top U hU⟩ instance : IsLocalAtSource P := by @@ -386,7 +391,7 @@ lemma isLocal_ringHomProperty_of_isLocalAtSource_of_isLocalAtTarget (R := CommRingCat.of S) (ι := s) (fun i : s ↦ (i : S)) (by simpa)).openCover intro i simp only [CommRingCat.coe_of, Set.setOf_mem_eq, id_eq, eq_mpr_eq_cast, - Scheme.AffineOpenCover.openCover_obj, Scheme.affineOpenCoverOfSpanRangeEqTop_obj, + Scheme.AffineOpenCover.openCover_obj, Scheme.affineOpenCoverOfSpanRangeEqTop_obj_carrier, Scheme.AffineOpenCover.openCover_map, Scheme.affineOpenCoverOfSpanRangeEqTop_map, ← Spec.map_comp] exact H i @@ -572,7 +577,7 @@ lemma iff_exists_appLE_locally (hQ : RingHom.StableUnderCompositionWithLocalizationAwaySource Q) (hQi : RespectsIso Q) [HasRingHomProperty P (Locally Q)] : P f ↔ ∀ (x : X), ∃ (U : Y.affineOpens) (V : X.affineOpens) (_ : x ∈ V.1) (e : V.1 ≤ f ⁻¹ᵁ U.1), - Q (f.appLE U V e) := by + Q (f.appLE U V e).hom := by have := respects_isOpenImmersion (P := P) (RingHom.locally_StableUnderCompositionWithLocalizationAwaySource hQ) refine ⟨fun hf x ↦ ?_, fun hf ↦ (IsLocalAtSource.iff_exists_resLE (P := P)).mpr <| fun x ↦ ?_⟩ @@ -602,7 +607,7 @@ lemma iff_exists_appLE_locally lemma iff_exists_appLE (hQ : StableUnderCompositionWithLocalizationAwaySource Q) : P f ↔ ∀ (x : X), ∃ (U : Y.affineOpens) (V : X.affineOpens) (_ : x ∈ V.1) (e : V.1 ≤ f ⁻¹ᵁ U.1), - Q (f.appLE U V e) := by + Q (f.appLE U V e).hom := by haveI inst : HasRingHomProperty P Q := inferInstance haveI : HasRingHomProperty P (Locally Q) := by apply @copy (P' := P) (Q := Q) (Q' := Locally Q) @@ -620,7 +625,7 @@ lemma locally_of_iff (hQl : LocalizationAwayPreserves Q) (hQa : StableUnderCompositionWithLocalizationAway Q) (h : ∀ {X Y : Scheme.{u}} (f : X ⟶ Y), P f ↔ ∀ (x : X), ∃ (U : Y.affineOpens) (V : X.affineOpens) (_ : x ∈ V.1) (e : V.1 ≤ f ⁻¹ᵁ U.1), - Q (f.appLE U V e)) : HasRingHomProperty P (Locally Q) where + Q (f.appLE U V e).hom) : HasRingHomProperty P (Locally Q) where isLocal_ringHomProperty := locally_propertyIsLocal hQl hQa eq_affineLocally' := by haveI : HasRingHomProperty (affineLocally (Locally Q)) (Locally Q) := diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean b/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean index c03365e68f13f..516521c7a6539 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean @@ -79,7 +79,7 @@ instance : IsLocalAtTarget @IsSeparated := by instance (R S : CommRingCat.{u}) (f : R ⟶ S) : IsSeparated (Spec.map f) := by constructor - letI := f.toAlgebra + letI := f.hom.toAlgebra show IsClosedImmersion (Limits.pullback.diagonal (Spec.map (CommRingCat.ofHom (algebraMap R S)))) rw [diagonal_Spec_map, MorphismProperty.cancel_right_of_respectsIso @IsClosedImmersion] exact .spec_of_surjective _ fun x ↦ ⟨.tmul R 1 x, diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Smooth.lean b/Mathlib/AlgebraicGeometry/Morphisms/Smooth.lean index 1d3998a469795..9d666644cb85d 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Smooth.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Smooth.lean @@ -60,7 +60,7 @@ standard smooth. @[mk_iff] class IsSmooth : Prop where exists_isStandardSmooth : ∀ (x : X), ∃ (U : Y.affineOpens) (V : X.affineOpens) (_ : x ∈ V.1) - (e : V.1 ≤ f ⁻¹ᵁ U.1), IsStandardSmooth.{0, 0} (f.appLE U V e) + (e : V.1 ≤ f ⁻¹ᵁ U.1), IsStandardSmooth.{0, 0} (f.appLE U V e).hom /-- The property of scheme morphisms `IsSmooth` is associated with the ring homomorphism property `Locally IsStandardSmooth.{0, 0}`. -/ @@ -97,7 +97,7 @@ standard smooth of relative dimension `n`. class IsSmoothOfRelativeDimension : Prop where exists_isStandardSmoothOfRelativeDimension : ∀ (x : X), ∃ (U : Y.affineOpens) (V : X.affineOpens) (_ : x ∈ V.1) (e : V.1 ≤ f ⁻¹ᵁ U.1), - IsStandardSmoothOfRelativeDimension.{0, 0} n (f.appLE U V e) + IsStandardSmoothOfRelativeDimension.{0, 0} n (f.appLE U V e).hom /-- If `f` is smooth of any relative dimension, it is smooth. -/ lemma IsSmoothOfRelativeDimension.isSmooth [IsSmoothOfRelativeDimension n f] : IsSmooth f where @@ -147,8 +147,10 @@ instance isSmoothOfRelativeDimension_comp {Z : Scheme.{u}} (g : Y ⟶ Z) have e : X.basicOpen s ≤ (f ≫ g) ⁻¹ᵁ U₂ := le_trans e₁ <| f.preimage_le_preimage_of_le <| le_trans (Y.basicOpen_le r) e₂ have heq : (f ≫ g).appLE U₂ (X.basicOpen s) e = g.appLE U₂ V₂ e₂ ≫ - algebraMap Γ(Y, V₂) Γ(Y, Y.basicOpen r) ≫ f.appLE (Y.basicOpen r) (X.basicOpen s) e₁ := by - rw [RingHom.algebraMap_toAlgebra, g.appLE_map_assoc, Scheme.appLE_comp_appLE] + CommRingCat.ofHom (algebraMap Γ(Y, V₂) Γ(Y, Y.basicOpen r)) ≫ + f.appLE (Y.basicOpen r) (X.basicOpen s) e₁ := by + rw [RingHom.algebraMap_toAlgebra, CommRingCat.ofHom_hom, + g.appLE_map_assoc, Scheme.appLE_comp_appLE] refine ⟨U₂, ⟨X.basicOpen s, V₁'.2.basicOpen s⟩, hx₁, e, heq ▸ ?_⟩ apply IsStandardSmoothOfRelativeDimension.comp ?_ hf₂ haveI : IsLocalization.Away r Γ(Y, Y.basicOpen r) := V₂.2.isLocalization_basicOpen r diff --git a/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean b/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean index 31a833ebabe90..a39842eb7658c 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean @@ -39,7 +39,7 @@ theorem Scheme.Hom.stalkMap_surjective (f : X.Hom Y) [SurjectiveOnStalks f] (x) namespace SurjectiveOnStalks instance (priority := 900) [IsOpenImmersion f] : SurjectiveOnStalks f := - ⟨fun _ ↦ (ConcreteCategory.bijective_of_isIso _).2⟩ + ⟨fun _ ↦ (ConcreteCategory.bijective_of_isIso (C := CommRingCat) _).2⟩ instance : MorphismProperty.IsMultiplicative @SurjectiveOnStalks where id_mem _ := inferInstance @@ -63,7 +63,7 @@ instance : IsLocalAtSource @SurjectiveOnStalks := eq_stalkwise ▸ stalkwise_isLocalAtSource_of_respectsIso RingHom.surjective_respectsIso lemma Spec_iff {R S : CommRingCat.{u}} {φ : R ⟶ S} : - SurjectiveOnStalks (Spec.map φ) ↔ RingHom.SurjectiveOnStalks φ := by + SurjectiveOnStalks (Spec.map φ) ↔ RingHom.SurjectiveOnStalks φ.hom := by rw [eq_stalkwise, stalkwise_Spec_map_iff RingHom.surjective_respectsIso, RingHom.SurjectiveOnStalks] @@ -72,7 +72,7 @@ instance : HasRingHomProperty @SurjectiveOnStalks RingHom.SurjectiveOnStalks := variable {f} in lemma iff_of_isAffine [IsAffine X] [IsAffine Y] : - SurjectiveOnStalks f ↔ RingHom.SurjectiveOnStalks (f.app ⊤) := by + SurjectiveOnStalks f ↔ RingHom.SurjectiveOnStalks (f.app ⊤).hom := by rw [← Spec_iff, MorphismProperty.arrow_mk_iso_iff @SurjectiveOnStalks (arrowIsoSpecΓOfIsAffine f)] theorem of_comp [SurjectiveOnStalks (f ≫ g)] : SurjectiveOnStalks f := by @@ -103,8 +103,7 @@ lemma isEmbedding_pullback {X Y S : Scheme.{u}} (f : X ⟶ S) (g : Y ⟶ S) [Sur .of_comp _ iS obtain ⟨φ, rfl⟩ : ∃ φ, Spec.map φ = f' := ⟨_, Spec.map_preimage _⟩ obtain ⟨ψ, rfl⟩ : ∃ ψ, Spec.map ψ = g' := ⟨_, Spec.map_preimage _⟩ - letI := φ.toAlgebra - letI := ψ.toAlgebra + algebraize [φ.hom, ψ.hom] rw [HasRingHomProperty.Spec_iff (P := @SurjectiveOnStalks)] at H convert ((iX.isOpenEmbedding.prodMap iY.isOpenEmbedding).isEmbedding.comp (PrimeSpectrum.isEmbedding_tensorProductTo_of_surjectiveOnStalks R A B H)).comp diff --git a/Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean b/Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean index d8cdaa1f4dd0a..32c4886d35bf2 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean @@ -105,7 +105,7 @@ lemma compactSpace_of_universallyClosed let 𝒰 : X.OpenCover := X.affineCover let U (i : 𝒰.J) : X.Opens := (𝒰.map i).opensRange let T : Scheme := Spec (.of <| MvPolynomial 𝒰.J K) - let q : T ⟶ Spec (.of K) := Spec.map MvPolynomial.C + let q : T ⟶ Spec (.of K) := Spec.map (CommRingCat.ofHom MvPolynomial.C) let Ti (i : 𝒰.J) : T.Opens := basicOpen (MvPolynomial.X i) let fT : pullback f q ⟶ T := pullback.snd f q let p : pullback f q ⟶ X := pullback.fst f q @@ -125,7 +125,7 @@ lemma compactSpace_of_universallyClosed let σ : Finset 𝒰.J := MvPolynomial.vars g let φ : MvPolynomial 𝒰.J K →+* MvPolynomial 𝒰.J K := (MvPolynomial.aeval fun i : 𝒰.J ↦ if i ∈ σ then MvPolynomial.X i else 0).toRingHom - let t' : T := (Spec.map φ).base t + let t' : T := (Spec.map (CommRingCat.ofHom φ)).base t have ht'g : t' ∈ PrimeSpectrum.basicOpen g := show φ g ∉ t.asIdeal from (show φ g = g from aeval_ite_mem_eq_self g subset_rfl).symm ▸ htU' have h : t' ∉ fT.base '' Z := hU'le ht'g diff --git a/Mathlib/AlgebraicGeometry/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/OpenImmersion.lean index 834a151635fb5..377bcde76e43c 100644 --- a/Mathlib/AlgebraicGeometry/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/OpenImmersion.lean @@ -177,12 +177,19 @@ theorem app_invApp' (U) (hU : U ≤ f.opensRange) : Y.presheaf.map (eqToHom (Opens.ext <| by simpa [Set.image_preimage_eq_inter_range])).op := PresheafedSpace.IsOpenImmersion.app_invApp _ _ -@[reassoc (attr := simp), elementwise (attr := simp)] +@[reassoc (attr := simp), elementwise nosimp] theorem appIso_inv_app (U) : (f.appIso U).inv ≫ f.app (f ''ᵁ U) = X.presheaf.map (eqToHom (preimage_image_eq f U)).op := (PresheafedSpace.IsOpenImmersion.invApp_app _ _).trans (by rw [eqToHom_op]) -@[reassoc (attr := simp), elementwise] +/-- +`elementwise` generates the `ConcreteCategory.instFunLike` lemma, we want `CommRingCat.Hom.hom`. +-/ +theorem appIso_inv_app_apply' (U) (x) : + f.app (f ''ᵁ U) ((f.appIso U).inv x) = X.presheaf.map (eqToHom (preimage_image_eq f U)).op x := + appIso_inv_app_apply f U x + +@[reassoc (attr := simp), elementwise nosimp] lemma appLE_appIso_inv {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion f] {U : Y.Opens} {V : X.Opens} (e : V ≤ f ⁻¹ᵁ U) : f.appLE U V e ≫ (f.appIso V).inv = @@ -231,8 +238,7 @@ lemma _root_.AlgebraicGeometry.IsOpenImmersion.of_isLocalization {R S} [CommRing IsOpenImmersion (Spec.map (CommRingCat.ofHom (algebraMap R S))) := by have e := (IsLocalization.algEquiv (.powers f) S (Localization.Away f)).symm.toAlgHom.comp_algebraMap - rw [← e, CommRingCat.ringHom_comp_eq_comp] - erw [Spec.map_comp] + rw [← e, CommRingCat.ofHom_comp, Spec.map_comp] have H : IsIso (CommRingCat.ofHom (IsLocalization.algEquiv (Submonoid.powers f) S (Localization.Away f)).symm.toAlgHom.toRingHom) := by exact inferInstanceAs (IsIso <| (IsLocalization.algEquiv @@ -678,7 +684,7 @@ theorem image_basicOpen {X Y : Scheme.{u}} (f : X ⟶ Y) [H : IsOpenImmersion f] (r : Γ(X, U)) : f ''ᵁ X.basicOpen r = Y.basicOpen ((f.appIso U).inv r) := by have e := Scheme.preimage_basicOpen f ((f.appIso U).inv r) - rw [Scheme.Hom.appIso_inv_app_apply, Scheme.basicOpen_res, inf_eq_right.mpr _] at e + rw [Scheme.Hom.appIso_inv_app_apply', Scheme.basicOpen_res, inf_eq_right.mpr _] at e · rw [← e, f.image_preimage_eq_opensRange_inter, inf_eq_right] refine Set.Subset.trans (Scheme.basicOpen_le _ _) (Set.image_subset_range _ _) · exact (X.basicOpen_le r).trans (f.preimage_image_eq _).ge diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean index 1ebe7c785f7bb..645435b9424b1 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean @@ -106,7 +106,7 @@ lemma basicOpenToSpec_app_top : noncomputable def toSpecZero : Proj 𝒜 ⟶ Spec (.of (𝒜 0)) := (Scheme.topIso _).inv ≫ (Scheme.isoOfEq _ (basicOpen_one _)).inv ≫ - basicOpenToSpec 𝒜 1 ≫ Spec.map (fromZeroRingHom 𝒜 _) + basicOpenToSpec 𝒜 1 ≫ Spec.map (CommRingCat.ofHom (fromZeroRingHom 𝒜 _)) variable {m} (f_deg : f ∈ 𝒜 m) (hm : 0 < m) @@ -177,8 +177,7 @@ lemma awayMap_awayToSection : apply Subtype.ext ext ⟨i, hi⟩ obtain ⟨⟨n, a, ⟨b, hb'⟩, i, rfl : _ = b⟩, rfl⟩ := mk_surjective a - simp only [CommRingCat.forget_obj, CommRingCat.coe_of, CommRingCat.ofHom, CommRingCat.coe_comp_of, - RingHom.coe_comp, Function.comp_apply, homOfLE_leOfHom] + simp only [homOfLE_leOfHom, CommRingCat.hom_comp, RingHom.coe_comp, Function.comp_apply] erw [ProjectiveSpectrum.Proj.awayToSection_apply] rw [val_awayMap_mk, Localization.mk_eq_mk', IsLocalization.map_mk', ← Localization.mk_eq_mk'] diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Proper.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Proper.lean index 6235852b32c97..511c2726b3d7f 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Proper.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Proper.lean @@ -106,7 +106,7 @@ instance isSeparated : IsSeparated (toSpecZero 𝒜) := by ← Spec.map_comp] erw [pullbackAwayιIso_inv_fst] congr 1 - ext x + ext x : 2 exact DFunLike.congr_fun (Algebra.TensorProduct.lift_comp_includeLeft (awayMapₐ 𝒜 j.2.2 rfl) (awayMapₐ 𝒜 i.2.2 (mul_comm _ _)) (fun _ _ ↦ .all _ _)).symm x · simp only [Iso.trans_hom, congrHom_hom, Category.assoc, Iso.hom_inv_id, Category.comp_id, @@ -115,7 +115,7 @@ instance isSeparated : IsSeparated (toSpecZero 𝒜) := by Spec.map_comp, e₂, e₁] erw [pullbackAwayιIso_inv_snd] congr 1 - ext x + ext x : 2 exact DFunLike.congr_fun (Algebra.TensorProduct.lift_comp_includeRight (awayMapₐ 𝒜 j.2.2 rfl) (awayMapₐ 𝒜 i.2.2 (mul_comm _ _)) (fun _ _ ↦ .all _ _)).symm x diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean index 781f0783739f0..38c08bfa1bb29 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean @@ -583,21 +583,25 @@ namespace ProjectiveSpectrum.Proj The ring map from `A⁰_ f` to the local sections of the structure sheaf of the projective spectrum of `A` on the basic open set `D(f)` defined by sending `s ∈ A⁰_f` to the section `x ↦ s` on `D(f)`. -/ -def awayToSection (f) : CommRingCat.of (A⁰_ f) ⟶ (structureSheaf 𝒜).1.obj (op (pbo f)) where - toFun s := - ⟨fun x ↦ HomogeneousLocalization.mapId 𝒜 (Submonoid.powers_le.mpr x.2) s, fun x ↦ by - obtain ⟨s, rfl⟩ := HomogeneousLocalization.mk_surjective s - obtain ⟨n, hn : f ^ n = s.den.1⟩ := s.den_mem - exact ⟨_, x.2, 𝟙 _, s.1, s.2, s.3, - fun x hsx ↦ x.2 (Ideal.IsPrime.mem_of_pow_mem inferInstance n (hn ▸ hsx)), fun _ ↦ rfl⟩⟩ - map_add' _ _ := by ext; simp only [map_add, HomogeneousLocalization.val_add, Proj.add_apply] - map_mul' _ _ := by ext; simp only [map_mul, HomogeneousLocalization.val_mul, Proj.mul_apply] - map_zero' := by ext; simp only [map_zero, HomogeneousLocalization.val_zero, Proj.zero_apply] - map_one' := by ext; simp only [map_one, HomogeneousLocalization.val_one, Proj.one_apply] +def awayToSection (f) : CommRingCat.of (A⁰_ f) ⟶ (structureSheaf 𝒜).1.obj (op (pbo f)) := + CommRingCat.ofHom + -- Have to hint `S`, otherwise it gets unfolded to `structureSheafInType` + -- causing `ext` to fail + (S := (structureSheaf 𝒜).1.obj (op (pbo f))) + { toFun s := + ⟨fun x ↦ HomogeneousLocalization.mapId 𝒜 (Submonoid.powers_le.mpr x.2) s, fun x ↦ by + obtain ⟨s, rfl⟩ := HomogeneousLocalization.mk_surjective s + obtain ⟨n, hn : f ^ n = s.den.1⟩ := s.den_mem + exact ⟨_, x.2, 𝟙 _, s.1, s.2, s.3, + fun x hsx ↦ x.2 (Ideal.IsPrime.mem_of_pow_mem inferInstance n (hn ▸ hsx)), fun _ ↦ rfl⟩⟩ + map_add' _ _ := by ext; simp only [map_add, HomogeneousLocalization.val_add, Proj.add_apply] + map_mul' _ _ := by ext; simp only [map_mul, HomogeneousLocalization.val_mul, Proj.mul_apply] + map_zero' := by ext; simp only [map_zero, HomogeneousLocalization.val_zero, Proj.zero_apply] + map_one' := by ext; simp only [map_one, HomogeneousLocalization.val_one, Proj.one_apply] } lemma awayToSection_germ (f x hx) : awayToSection 𝒜 f ≫ (structureSheaf 𝒜).presheaf.germ _ x hx = - (HomogeneousLocalization.mapId 𝒜 (Submonoid.powers_le.mpr hx)) ≫ + CommRingCat.ofHom (HomogeneousLocalization.mapId 𝒜 (Submonoid.powers_le.mpr hx)) ≫ (Proj.stalkIso' 𝒜 x).toCommRingCatIso.inv := by ext z apply (Proj.stalkIso' 𝒜 x).eq_symm_apply.mpr @@ -625,7 +629,7 @@ def awayToΓ (f) : CommRingCat.of (A⁰_ f) ⟶ LocallyRingedSpace.Γ.obj (op <| lemma awayToΓ_ΓToStalk (f) (x) : awayToΓ 𝒜 f ≫ (Proj| pbo f).presheaf.Γgerm x = - HomogeneousLocalization.mapId 𝒜 (Submonoid.powers_le.mpr x.2) ≫ + CommRingCat.ofHom (HomogeneousLocalization.mapId 𝒜 (Submonoid.powers_le.mpr x.2)) ≫ (Proj.stalkIso' 𝒜 x.1).toCommRingCatIso.inv ≫ ((Proj.toLocallyRingedSpace 𝒜).restrictStalkIso (Opens.isOpenEmbedding _) x).inv := by rw [awayToΓ, Category.assoc, ← Category.assoc _ (Iso.inv _), @@ -648,9 +652,9 @@ open HomogeneousLocalization IsLocalRing lemma toSpec_base_apply_eq_comap {f} (x : Proj| pbo f) : (toSpec 𝒜 f).base x = PrimeSpectrum.comap (mapId 𝒜 (Submonoid.powers_le.mpr x.2)) (closedPoint (AtPrime 𝒜 x.1.asHomogeneousIdeal.toIdeal)) := by - show PrimeSpectrum.comap (awayToΓ 𝒜 f ≫ (Proj| pbo f).presheaf.Γgerm x) + show PrimeSpectrum.comap (awayToΓ 𝒜 f ≫ (Proj| pbo f).presheaf.Γgerm x).hom (IsLocalRing.closedPoint ((Proj| pbo f).presheaf.stalk x)) = _ - rw [awayToΓ_ΓToStalk, CommRingCat.comp_eq_ring_hom_comp, PrimeSpectrum.comap_comp] + rw [awayToΓ_ΓToStalk, CommRingCat.hom_comp, PrimeSpectrum.comap_comp] exact congr(PrimeSpectrum.comap _ $(@IsLocalRing.comap_closedPoint (HomogeneousLocalization.AtPrime 𝒜 x.1.asHomogeneousIdeal.toIdeal) _ _ ((Proj| pbo f).presheaf.stalk x) _ _ _ (isLocalHom_of_isIso _))) @@ -696,7 +700,7 @@ lemma toStalk_stalkMap_toSpec (f) (x) : StructureSheaf.toStalk _ _ ≫ (toSpec 𝒜 f).stalkMap x = awayToΓ 𝒜 f ≫ (Proj| pbo f).presheaf.Γgerm x := by rw [StructureSheaf.toStalk, Category.assoc] - simp_rw [CommRingCat.coe_of, ← Spec.locallyRingedSpaceObj_presheaf'] + simp_rw [← Spec.locallyRingedSpaceObj_presheaf'] rw [LocallyRingedSpace.stalkMap_germ (toSpec 𝒜 f), toOpen_toSpec_val_c_app_assoc, Presheaf.germ_res] rfl @@ -779,11 +783,11 @@ def specStalkEquiv (f) (x : pbo f) {m} (f_deg : f ∈ 𝒜 m) (hm : 0 < m) : lemma toStalk_specStalkEquiv (f) (x : pbo f) {m} (f_deg : f ∈ 𝒜 m) (hm : 0 < m) : StructureSheaf.toStalk (A⁰_ f) ((toSpec 𝒜 f).base x) ≫ (specStalkEquiv 𝒜 f x f_deg hm).hom = - (mapId _ <| Submonoid.powers_le.mpr x.2 : (A⁰_ f) →+* AtPrime 𝒜 x.1.1.toIdeal) := + CommRingCat.ofHom (mapId _ <| Submonoid.powers_le.mpr x.2) := letI : Algebra (Away 𝒜 f) (AtPrime 𝒜 x.1.asHomogeneousIdeal.toIdeal) := (mapId 𝒜 (Submonoid.powers_le.mpr x.2)).toAlgebra letI := isLocalization_atPrime 𝒜 f x f_deg hm - (IsLocalization.algEquiv + CommRingCat.hom_ext (IsLocalization.algEquiv (R := A⁰_ f) (M := ((toSpec 𝒜 f).base x).asIdeal.primeCompl) (S := (Spec.structureSheaf (A⁰_ f)).presheaf.stalk ((toSpec 𝒜 f).base x)) @@ -793,10 +797,12 @@ lemma stalkMap_toSpec (f) (x : pbo f) {m} (f_deg : f ∈ 𝒜 m) (hm : 0 < m) : (toSpec 𝒜 f).stalkMap x = (specStalkEquiv 𝒜 f x f_deg hm).hom ≫ (Proj.stalkIso' 𝒜 x.1).toCommRingCatIso.inv ≫ ((Proj.toLocallyRingedSpace 𝒜).restrictStalkIso (Opens.isOpenEmbedding _) x).inv := - IsLocalization.ringHom_ext (R := A⁰_ f) ((toSpec 𝒜 f).base x).asIdeal.primeCompl - (S := (Spec.structureSheaf (A⁰_ f)).presheaf.stalk ((toSpec 𝒜 f).base x)) <| - (toStalk_stalkMap_toSpec _ _ _).trans <| by - rw [awayToΓ_ΓToStalk, ← toStalk_specStalkEquiv, Category.assoc]; rfl + CommRingCat.hom_ext <| + IsLocalization.ringHom_ext (R := A⁰_ f) ((toSpec 𝒜 f).base x).asIdeal.primeCompl + (S := (Spec.structureSheaf (A⁰_ f)).presheaf.stalk ((toSpec 𝒜 f).base x)) <| + CommRingCat.hom_ext_iff.mp <| + (toStalk_stalkMap_toSpec _ _ _).trans <| by + rw [awayToΓ_ΓToStalk, ← toStalk_specStalkEquiv, Category.assoc]; rfl lemma isIso_toSpec (f) {m} (f_deg : f ∈ 𝒜 m) (hm : 0 < m) : IsIso (toSpec 𝒜 f) := by diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/StructureSheaf.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/StructureSheaf.lean index 1b18003cb9d51..7532e95080f90 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/StructureSheaf.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/StructureSheaf.lean @@ -177,7 +177,7 @@ structure presheaf. -/ @[simps] def structurePresheafInCommRing : Presheaf CommRingCat (ProjectiveSpectrum.top 𝒜) where obj U := CommRingCat.of ((structureSheafInType 𝒜).1.obj U) - map i := + map i := CommRingCat.ofHom { toFun := (structureSheafInType 𝒜).1.map i map_zero' := rfl map_add' := fun _ _ => rfl @@ -186,7 +186,7 @@ def structurePresheafInCommRing : Presheaf CommRingCat (ProjectiveSpectrum.top -- These lemmas have always been bad (https://github.com/leanprover-community/mathlib4/issues/7657), but https://github.com/leanprover/lean4/pull/2644 made `simp` start noticing attribute [nolint simpNF] - AlgebraicGeometry.ProjectiveSpectrum.StructureSheaf.structurePresheafInCommRing_map_apply + AlgebraicGeometry.ProjectiveSpectrum.StructureSheaf.structurePresheafInCommRing_map_hom_apply /-- Some glue, verifying that the structure presheaf valued in `CommRing` agrees with the `Type` valued structure presheaf. -/ @@ -241,12 +241,13 @@ def Proj.toSheafedSpace : SheafedSpace CommRingCat where implemented as a subtype of dependent functions to localizations at homogeneous prime ideals, and evaluates the section on the point corresponding to a given homogeneous prime ideal. -/ def openToLocalization (U : Opens (ProjectiveSpectrum.top 𝒜)) (x : ProjectiveSpectrum.top 𝒜) - (hx : x ∈ U) : (Proj.structureSheaf 𝒜).1.obj (op U) ⟶ CommRingCat.of (at x) where - toFun s := (s.1 ⟨x, hx⟩ : _) - map_one' := rfl - map_mul' _ _ := rfl - map_zero' := rfl - map_add' _ _ := rfl + (hx : x ∈ U) : (Proj.structureSheaf 𝒜).1.obj (op U) ⟶ CommRingCat.of (at x) := + CommRingCat.ofHom + { toFun s := (s.1 ⟨x, hx⟩ : _) + map_one' := rfl + map_mul' _ _ := rfl + map_zero' := rfl + map_add' _ _ := rfl } /-- The ring homomorphism from the stalk of the structure sheaf of `Proj` at a point corresponding to a homogeneous prime ideal `x` to the *homogeneous localization* at `x`, @@ -270,7 +271,7 @@ theorem germ_comp_stalkToFiberRingHom theorem stalkToFiberRingHom_germ (U : Opens (ProjectiveSpectrum.top 𝒜)) (x : ProjectiveSpectrum.top 𝒜) (hx : x ∈ U) (s : (Proj.structureSheaf 𝒜).1.obj (op U)) : stalkToFiberRingHom 𝒜 x ((Proj.structureSheaf 𝒜).presheaf.germ _ x hx s) = s.1 ⟨x, hx⟩ := - RingHom.ext_iff.1 (germ_comp_stalkToFiberRingHom 𝒜 U x hx) s + RingHom.ext_iff.1 (CommRingCat.hom_ext_iff.mp (germ_comp_stalkToFiberRingHom 𝒜 U x hx)) s @[deprecated (since := "2024-07-30")] alias stalkToFiberRingHom_germ' := stalkToFiberRingHom_germ @@ -308,6 +309,11 @@ def homogeneousLocalizationToStalk (x : ProjectiveSpectrum.top 𝒜) (y : at x) ProjectiveSpectrum.basicOpen 𝒜 g.den.1 ⊓ ProjectiveSpectrum.basicOpen 𝒜 c) ⟨⟨mem_basicOpen_den _ x f, mem_basicOpen_den _ x g⟩, hc⟩ (homOfLE inf_le_left ≫ homOfLE inf_le_left) (homOfLE inf_le_left ≫ homOfLE inf_le_right) + -- Go from `ConcreteCategory.instFunLike` to `CommRingCat.Hom.hom` + show (Proj.structureSheaf 𝒜).presheaf.map (homOfLE inf_le_left ≫ homOfLE inf_le_left).op + (sectionInBasicOpen 𝒜 x f) = + (Proj.structureSheaf 𝒜).presheaf.map (homOfLE inf_le_left ≫ homOfLE inf_le_right).op + (sectionInBasicOpen 𝒜 x g) apply Subtype.ext ext ⟨t, ⟨htf, htg⟩, ht'⟩ rw [Proj.res_apply, Proj.res_apply] @@ -320,10 +326,15 @@ def homogeneousLocalizationToStalk (x : ProjectiveSpectrum.top 𝒜) (y : at x) lemma homogeneousLocalizationToStalk_stalkToFiberRingHom (x z) : homogeneousLocalizationToStalk 𝒜 x (stalkToFiberRingHom 𝒜 x z) = z := by obtain ⟨U, hxU, s, rfl⟩ := (Proj.structureSheaf 𝒜).presheaf.germ_exist x z + show homogeneousLocalizationToStalk 𝒜 x ((stalkToFiberRingHom 𝒜 x).hom + (((Proj.structureSheaf 𝒜).presheaf.germ U x hxU) s)) = + ((Proj.structureSheaf 𝒜).presheaf.germ U x hxU) s obtain ⟨V, hxV, i, n, a, b, h, e⟩ := s.2 ⟨x, hxU⟩ simp only at e rw [stalkToFiberRingHom_germ, homogeneousLocalizationToStalk, e ⟨x, hxV⟩, Quotient.liftOn'_mk''] - refine Presheaf.germ_ext _ V hxV (by exact homOfLE <| fun _ h' ↦ h ⟨_, h'⟩) i ?_ + refine Presheaf.germ_ext (C := CommRingCat) _ V hxV (homOfLE <| fun _ h' ↦ h ⟨_, h'⟩) i ?_ + change ((Proj.structureSheaf 𝒜).presheaf.map (homOfLE <| fun _ h' ↦ h ⟨_, h'⟩).op) _ = + ((Proj.structureSheaf 𝒜).presheaf.map i.op) s apply Subtype.ext ext ⟨t, ht⟩ rw [Proj.res_apply, Proj.res_apply] @@ -340,7 +351,7 @@ lemma stalkToFiberRingHom_homogeneousLocalizationToStalk (x z) : and homogeneous localization at `x` for any point `x` in `Proj`. -/ def Proj.stalkIso' (x : ProjectiveSpectrum.top 𝒜) : (Proj.structureSheaf 𝒜).presheaf.stalk x ≃+* at x where - __ := stalkToFiberRingHom _ x + __ := (stalkToFiberRingHom _ x).hom invFun := homogeneousLocalizationToStalk 𝒜 x left_inv := homogeneousLocalizationToStalk_stalkToFiberRingHom 𝒜 x right_inv := stalkToFiberRingHom_homogeneousLocalizationToStalk 𝒜 x diff --git a/Mathlib/AlgebraicGeometry/Properties.lean b/Mathlib/AlgebraicGeometry/Properties.lean index 79091bf28dd7a..a009eb29b44fc 100644 --- a/Mathlib/AlgebraicGeometry/Properties.lean +++ b/Mathlib/AlgebraicGeometry/Properties.lean @@ -56,6 +56,7 @@ theorem isReduced_of_isReduced_stalk [∀ x : X, _root_.IsReduced (X.presheaf.st refine ⟨fun U => ⟨fun s hs => ?_⟩⟩ apply Presheaf.section_ext X.sheaf U s 0 intro x hx + show (X.sheaf.presheaf.germ U x hx) s = (X.sheaf.presheaf.germ U x hx) 0 rw [RingHom.map_zero] change X.presheaf.germ U x hx s = 0 exact (hs.map _).eq_zero @@ -64,12 +65,13 @@ instance isReduced_stalk_of_isReduced [IsReduced X] (x : X) : _root_.IsReduced (X.presheaf.stalk x) := by constructor rintro g ⟨n, e⟩ - obtain ⟨U, hxU, s, rfl⟩ := X.presheaf.germ_exist x g - rw [← map_pow, ← map_zero (X.presheaf.germ _ x hxU)] at e - obtain ⟨V, hxV, iU, iV, e'⟩ := X.presheaf.germ_eq x hxU hxU _ 0 e + obtain ⟨U, hxU, s, (rfl : (X.presheaf.germ U x hxU) s = g)⟩ := X.presheaf.germ_exist x g + rw [← map_pow, ← map_zero (X.presheaf.germ _ x hxU).hom] at e + obtain ⟨V, hxV, iU, iV, (e' : (X.presheaf.map iU.op) (s ^ n) = (X.presheaf.map iV.op) 0)⟩ := + X.presheaf.germ_eq x hxU hxU _ 0 e rw [map_pow, map_zero] at e' replace e' := (IsNilpotent.mk _ _ e').eq_zero (R := Γ(X, V)) - rw [← X.presheaf.germ_res iU x hxV, comp_apply, e', map_zero] + rw [← X.presheaf.germ_res iU x hxV, CommRingCat.comp_apply, e', map_zero] theorem isReduced_of_isOpenImmersion {X Y : Scheme} (f : X ⟶ Y) [H : IsOpenImmersion f] [IsReduced Y] : IsReduced X := by @@ -78,7 +80,7 @@ theorem isReduced_of_isOpenImmersion {X Y : Scheme} (f : X ⟶ Y) [H : IsOpenImm have : U = f ⁻¹ᵁ f ''ᵁ U := by ext1; exact (Set.preimage_image_eq _ H.base_open.injective).symm rw [this] - exact isReduced_of_injective (inv <| f.app (f ''ᵁ U)) + exact isReduced_of_injective (inv <| f.app (f ''ᵁ U)).hom (asIso <| f.app (f ''ᵁ U) : Γ(Y, f ''ᵁ U) ≅ _).symm.commRingCatIsoToRingEquiv.injective instance {R : CommRingCat.{u}} [H : _root_.IsReduced R] : IsReduced (Spec R) := by @@ -86,14 +88,14 @@ instance {R : CommRingCat.{u}} [H : _root_.IsReduced R] : IsReduced (Spec R) := intro x; dsimp have : _root_.IsReduced (CommRingCat.of <| Localization.AtPrime (PrimeSpectrum.asIdeal x)) := by dsimp; infer_instance - exact isReduced_of_injective (StructureSheaf.stalkIso R x).hom + exact isReduced_of_injective (StructureSheaf.stalkIso R x).hom.hom (StructureSheaf.stalkIso R x).commRingCatIsoToRingEquiv.injective theorem affine_isReduced_iff (R : CommRingCat) : IsReduced (Spec R) ↔ _root_.IsReduced R := by refine ⟨?_, fun h => inferInstance⟩ intro h - exact isReduced_of_injective (Scheme.ΓSpecIso R).inv + exact isReduced_of_injective (Scheme.ΓSpecIso R).inv.hom (Scheme.ΓSpecIso R).symm.commRingCatIsoToRingEquiv.injective theorem isReduced_of_isAffine_isReduced [IsAffine X] [_root_.IsReduced Γ(X, ⊤)] : @@ -139,15 +141,16 @@ theorem eq_zero_of_basicOpen_eq_bot {X : Scheme} [hX : IsReduced X] {U : X.Opens (s : Γ(X, U)) (hs : X.basicOpen s = ⊥) : s = 0 := by apply TopCat.Presheaf.section_ext X.sheaf U intro x hx + show (X.sheaf.presheaf.germ U x hx) s = (X.sheaf.presheaf.germ U x hx) 0 rw [RingHom.map_zero] - show X.presheaf.germ U x hx s = 0 induction U using reduce_to_affine_global generalizing hX with | h₁ X U H => obtain ⟨V, hx, i, H⟩ := H ⟨x, hx⟩ specialize H (X.presheaf.map i.op s) rw [Scheme.basicOpen_res, hs] at H specialize H (inf_bot_eq _) x hx - rw [TopCat.Presheaf.germ_res_apply] at H + -- This seems to be related to a mismatch of `X.sheaf.presheaf` and `X.presheaf` in `H` + rw [← CommRingCat.germ_res_apply X.sheaf.presheaf i x hx s] exact H | h₂ X Y f => refine ⟨f ⁻¹ᵁ f.opensRange, f.opensRange, by ext1; simp, rfl, ?_⟩ @@ -155,14 +158,15 @@ theorem eq_zero_of_basicOpen_eq_bot {X : Scheme} [hX : IsReduced X] {U : X.Opens haveI := isReduced_of_isOpenImmersion f specialize H (f.app _ s) _ x ⟨x, rfl⟩ · rw [← Scheme.preimage_basicOpen, hs]; ext1; simp [Opens.map] - · rw [← Scheme.stalkMap_germ_apply f ⟨_, _⟩ x] at H + · have H : (X.presheaf.germ _ x _).hom _ = 0 := H + rw [← Scheme.stalkMap_germ_apply f ⟨_, _⟩ x] at H apply_fun inv <| f.stalkMap x at H - rw [CategoryTheory.IsIso.hom_inv_id_apply, map_zero] at H + rw [← CommRingCat.comp_apply, CategoryTheory.IsIso.hom_inv_id, map_zero] at H exact H | h₃ R => rw [basicOpen_eq_of_affine', PrimeSpectrum.basicOpen_eq_bot_iff] at hs - replace hs := (hs.map (Scheme.ΓSpecIso R).inv).eq_zero - rw [Iso.hom_inv_id_apply] at hs + replace hs := (hs.map (Scheme.ΓSpecIso R).inv.hom).eq_zero + rw [← CommRingCat.comp_apply, Iso.hom_inv_id, CommRingCat.id_apply] at hs rw [hs, map_zero] @[simp] diff --git a/Mathlib/AlgebraicGeometry/ResidueField.lean b/Mathlib/AlgebraicGeometry/ResidueField.lean index 6ce1ecd8378e6..dc2f20d9029f1 100644 --- a/Mathlib/AlgebraicGeometry/ResidueField.lean +++ b/Mathlib/AlgebraicGeometry/ResidueField.lean @@ -46,7 +46,7 @@ instance (x : X) : Field (X.residueField x) := /-- The residue map from the stalk to the residue field. -/ def residue (X : Scheme.{u}) (x) : X.presheaf.stalk x ⟶ X.residueField x := - IsLocalRing.residue _ + CommRingCat.ofHom (IsLocalRing.residue (X.presheaf.stalk x)) /-- See `AlgebraicGeometry.IsClosedImmersion.Spec_map_residue` for the stronger result that `Spec.map (X.residue x)` is a closed immersion. -/ @@ -70,15 +70,15 @@ instance (X : Scheme.{u}) (x) : Epi (X.residue x) := /-- If `K` is a field and `f : 𝒪_{X, x} ⟶ K` is a ring map, then this is the induced map `κ(x) ⟶ K`. -/ def descResidueField {K : Type u} [Field K] {X : Scheme.{u}} {x : X} - (f : X.presheaf.stalk x ⟶ .of K) [IsLocalHom f] : + (f : X.presheaf.stalk x ⟶ .of K) [IsLocalHom f.hom] : X.residueField x ⟶ .of K := - IsLocalRing.ResidueField.lift (S := K) f + CommRingCat.ofHom (IsLocalRing.ResidueField.lift (S := K) f.hom) @[reassoc (attr := simp)] lemma residue_descResidueField {K : Type u} [Field K] {X : Scheme.{u}} {x} - (f : X.presheaf.stalk x ⟶ .of K) [IsLocalHom f] : + (f : X.presheaf.stalk x ⟶ .of K) [IsLocalHom f.hom] : X.residue x ≫ X.descResidueField f = f := - RingHom.ext fun _ ↦ rfl + CommRingCat.hom_ext <| RingHom.ext fun _ ↦ rfl /-- If `U` is an open of `X` containing `x`, we have a canonical ring map from the sections @@ -116,7 +116,7 @@ variable {X Y : Scheme.{u}} (f : X ⟶ Y) a morphism of residue fields in the other direction. -/ def Hom.residueFieldMap (f : X.Hom Y) (x : X) : Y.residueField (f.base x) ⟶ X.residueField x := - IsLocalRing.ResidueField.map (f.stalkMap x) + CommRingCat.ofHom <| IsLocalRing.ResidueField.map (f.stalkMap x).hom @[reassoc] lemma residue_residueFieldMap (x : X) : @@ -259,16 +259,17 @@ lemma range_fromSpecResidueField (x : X.carrier) : exact ⟨closedPoint (X.residueField x), rfl⟩ lemma descResidueField_fromSpecResidueField {K : Type*} [Field K] (X : Scheme) {x} - (f : X.presheaf.stalk x ⟶ .of K) [IsLocalHom f] : + (f : X.presheaf.stalk x ⟶ .of K) [IsLocalHom f.hom] : Spec.map (X.descResidueField f) ≫ X.fromSpecResidueField x = Spec.map f ≫ X.fromSpecStalk x := by simp [fromSpecResidueField, ← Spec.map_comp_assoc] lemma descResidueField_stalkClosedPointTo_fromSpecResidueField (K : Type u) [Field K] (X : Scheme.{u}) (f : Spec (.of K) ⟶ X) : - Spec.map (X.descResidueField (Scheme.stalkClosedPointTo f)) ≫ + Spec.map (@descResidueField (CommRingCat.of K) _ X _ (Scheme.stalkClosedPointTo f) + _) ≫ X.fromSpecResidueField (f.base (closedPoint K)) = f := by - erw [X.descResidueField_fromSpecResidueField] + rw [X.descResidueField_fromSpecResidueField] rw [Scheme.Spec_stalkClosedPointTo_fromSpecStalk] end fromResidueField diff --git a/Mathlib/AlgebraicGeometry/Restrict.lean b/Mathlib/AlgebraicGeometry/Restrict.lean index 1d336822caba3..b0cb9b5db318c 100644 --- a/Mathlib/AlgebraicGeometry/Restrict.lean +++ b/Mathlib/AlgebraicGeometry/Restrict.lean @@ -175,15 +175,15 @@ def opensRestrict : instance ΓRestrictAlgebra {X : Scheme.{u}} (U : X.Opens) : Algebra (Γ(X, ⊤)) Γ(U, ⊤) := - U.ι.appTop.toAlgebra + U.ι.appTop.hom.toAlgebra lemma Scheme.map_basicOpen (r : Γ(U, ⊤)) : U.ι ''ᵁ U.toScheme.basicOpen r = X.basicOpen (X.presheaf.map (eqToHom U.isOpenEmbedding_obj_top.symm).op r) := by refine (Scheme.image_basicOpen (X.ofRestrict U.isOpenEmbedding) r).trans ?_ rw [← Scheme.basicOpen_res_eq _ _ (eqToHom U.isOpenEmbedding_obj_top).op] - rw [← comp_apply, ← CategoryTheory.Functor.map_comp, ← op_comp, eqToHom_trans, eqToHom_refl, - op_id, CategoryTheory.Functor.map_id] + rw [← CommRingCat.comp_apply, ← CategoryTheory.Functor.map_comp, ← op_comp, eqToHom_trans, + eqToHom_refl, op_id, CategoryTheory.Functor.map_id] congr exact PresheafedSpace.IsOpenImmersion.ofRestrict_invApp _ _ _ @@ -617,7 +617,7 @@ def morphismRestrictRestrictBasicOpen {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Op have e := Scheme.preimage_basicOpen U.ι r rw [Scheme.Opens.ι_app] at e rw [← U.toScheme.basicOpen_res_eq _ (eqToHom U.inclusion'_map_eq_top).op] - erw [← comp_apply] + erw [← CommRingCat.comp_apply] erw [← Y.presheaf.map_comp] rw [eqToHom_op, eqToHom_op, eqToHom_map, eqToHom_trans] erw [← e] diff --git a/Mathlib/AlgebraicGeometry/Scheme.lean b/Mathlib/AlgebraicGeometry/Scheme.lean index dc282d80303ad..5fffbe52e634b 100644 --- a/Mathlib/AlgebraicGeometry/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/Scheme.lean @@ -155,7 +155,7 @@ lemma appLE_eq_app {U : Y.Opens} : (app_eq_appLE f).symm lemma appLE_congr (e : V ≤ f ⁻¹ᵁ U) (e₁ : U = U') (e₂ : V = V') - (P : ∀ {R S : Type u} [CommRing R] [CommRing S] (_ : R →+* S), Prop) : + (P : ∀ {R S : CommRingCat.{u}} (_ : R ⟶ S), Prop) : P (f.appLE U V e) ↔ P (f.appLE U' V' (e₁ ▸ e₂ ▸ e)) := by subst e₁; subst e₂; rfl @@ -410,14 +410,15 @@ variable {R S : CommRingCat.{u}} (f : R ⟶ S) lemma Spec_carrier (R : CommRingCat.{u}) : (Spec R).carrier = PrimeSpectrum R := rfl lemma Spec_sheaf (R : CommRingCat.{u}) : (Spec R).sheaf = Spec.structureSheaf R := rfl lemma Spec_presheaf (R : CommRingCat.{u}) : (Spec R).presheaf = (Spec.structureSheaf R).1 := rfl -lemma Spec.map_base : (Spec.map f).base = PrimeSpectrum.comap f := rfl -lemma Spec.map_base_apply (x : Spec S) : (Spec.map f).base x = PrimeSpectrum.comap f x := rfl +lemma Spec.map_base : (Spec.map f).base = PrimeSpectrum.comap f.hom := rfl +lemma Spec.map_base_apply (x : Spec S) : (Spec.map f).base x = PrimeSpectrum.comap f.hom x := rfl lemma Spec.map_app (U) : - (Spec.map f).app U = StructureSheaf.comap f U (Spec.map f ⁻¹ᵁ U) le_rfl := rfl + (Spec.map f).app U = + CommRingCat.ofHom (StructureSheaf.comap f.hom U (Spec.map f ⁻¹ᵁ U) le_rfl) := rfl lemma Spec.map_appLE {U V} (e : U ≤ Spec.map f ⁻¹ᵁ V) : - (Spec.map f).appLE V U e = StructureSheaf.comap f V U e := rfl + (Spec.map f).appLE V U e = CommRingCat.ofHom (StructureSheaf.comap f.hom V U e) := rfl instance {A : CommRingCat} [Nontrivial A] : Nonempty (Spec A) := inferInstanceAs <| Nonempty (PrimeSpectrum A) @@ -547,7 +548,8 @@ theorem basicOpen_le : X.basicOpen f ≤ U := @[sheaf_restrict] lemma basicOpen_restrict (i : V ⟶ U) (f : Γ(X, U)) : - X.basicOpen (f |_ₕ i) ≤ X.basicOpen f := + -- Help `restrict` to infer which forgetful functor we're taking + X.basicOpen (TopCat.Presheaf.restrict (C := CommRingCat) f i) ≤ X.basicOpen f := (Scheme.basicOpen_res _ _ _).trans_le inf_le_right @[simp] @@ -561,7 +563,7 @@ theorem preimage_basicOpen_top {X Y : Scheme.{u}} (f : X ⟶ Y) (r : Γ(Y, ⊤)) lemma basicOpen_appLE {X Y : Scheme.{u}} (f : X ⟶ Y) (U : X.Opens) (V : Y.Opens) (e : U ≤ f ⁻¹ᵁ V) (s : Γ(Y, V)) : X.basicOpen (f.appLE V U e s) = U ⊓ f ⁻¹ᵁ (Y.basicOpen s) := by - simp only [preimage_basicOpen, Hom.appLE, CommRingCat.coe_comp_of, RingHom.coe_comp, + simp only [preimage_basicOpen, Hom.appLE, CommRingCat.comp_apply, RingHom.coe_comp, Function.comp_apply] rw [basicOpen_res] @@ -581,7 +583,7 @@ theorem basicOpen_of_isUnit {f : Γ(X, U)} (hf : IsUnit f) : X.basicOpen f = U : instance algebra_section_section_basicOpen {X : Scheme} {U : X.Opens} (f : Γ(X, U)) : Algebra Γ(X, U) Γ(X, X.basicOpen f) := - (X.presheaf.map (homOfLE <| X.basicOpen_le f : _ ⟶ U).op).toAlgebra + (X.presheaf.map (homOfLE <| X.basicOpen_le f : _ ⟶ U).op).hom.toAlgebra end BasicOpen @@ -626,7 +628,7 @@ theorem basicOpen_eq_of_affine {R : CommRingCat} (f : R) : ext x simp only [SetLike.mem_coe, Scheme.mem_basicOpen_top, Opens.coe_top] suffices IsUnit (StructureSheaf.toStalk R x f) ↔ f ∉ PrimeSpectrum.asIdeal x by exact this - rw [← isUnit_map_iff (StructureSheaf.stalkToFiberRingHom R x), + rw [← isUnit_map_iff (StructureSheaf.stalkToFiberRingHom R x).hom, StructureSheaf.stalkToFiberRingHom_toStalk] exact (IsLocalization.AtPrime.isUnit_to_map_iff (Localization.AtPrime (PrimeSpectrum.asIdeal x)) @@ -683,7 +685,7 @@ namespace Scheme variable {X Y : Scheme.{u}} (f : X ⟶ Y) -instance (x) : IsLocalHom (f.stalkMap x) := +instance (x) : IsLocalHom (f.stalkMap x).hom := f.prop x @[simp] @@ -704,7 +706,7 @@ lemma stalkSpecializes_stalkMap (x x' : X) lemma stalkSpecializes_stalkMap_apply (x x' : X) (h : x ⤳ x') (y) : f.stalkMap x (Y.presheaf.stalkSpecializes (f.base.map_specializes h) y) = (X.presheaf.stalkSpecializes h (f.stalkMap x' y)) := - DFunLike.congr_fun (stalkSpecializes_stalkMap f x x' h) y + DFunLike.congr_fun (CommRingCat.hom_ext_iff.mp (stalkSpecializes_stalkMap f x x' h)) y @[reassoc] lemma stalkMap_congr (f g : X ⟶ Y) (hfg : f = g) (x x' : X) @@ -733,7 +735,7 @@ lemma stalkMap_hom_inv (e : X ≅ Y) (y : Y) : lemma stalkMap_hom_inv_apply (e : X ≅ Y) (y : Y) (z) : e.inv.stalkMap y (e.hom.stalkMap (e.inv.base y) z) = (Y.presheaf.stalkCongr (.of_eq (by simp))).hom z := - DFunLike.congr_fun (stalkMap_hom_inv e y) z + DFunLike.congr_fun (CommRingCat.hom_ext_iff.mp (stalkMap_hom_inv e y)) z @[reassoc (attr := simp)] lemma stalkMap_inv_hom (e : X ≅ Y) (x : X) : @@ -745,7 +747,7 @@ lemma stalkMap_inv_hom (e : X ≅ Y) (x : X) : lemma stalkMap_inv_hom_apply (e : X ≅ Y) (x : X) (y) : e.hom.stalkMap x (e.inv.stalkMap (e.hom.base x) y) = (X.presheaf.stalkCongr (.of_eq (by simp))).hom y := - DFunLike.congr_fun (stalkMap_inv_hom e x) y + DFunLike.congr_fun (CommRingCat.hom_ext_iff.mp (stalkMap_inv_hom e x)) y @[reassoc (attr := simp)] lemma stalkMap_germ (U : Y.Opens) (x : X) (hx : f.base x ∈ U) : @@ -769,8 +771,8 @@ open IsLocalRing @[simp] lemma Spec_closedPoint {R S : CommRingCat} [IsLocalRing R] [IsLocalRing S] - {f : R ⟶ S} [IsLocalHom f] : (Spec.map f).base (closedPoint S) = closedPoint R := - IsLocalRing.comap_closedPoint f + {f : R ⟶ S} [IsLocalHom f.hom] : (Spec.map f).base (closedPoint S) = closedPoint R := + IsLocalRing.comap_closedPoint f.hom end IsLocalRing diff --git a/Mathlib/AlgebraicGeometry/Spec.lean b/Mathlib/AlgebraicGeometry/Spec.lean index b16f662644390..264ee6873f818 100644 --- a/Mathlib/AlgebraicGeometry/Spec.lean +++ b/Mathlib/AlgebraicGeometry/Spec.lean @@ -60,7 +60,7 @@ def Spec.topObj (R : CommRingCat.{u}) : TopCat := /-- The induced map of a ring homomorphism on the ring spectra, as a morphism of topological spaces. -/ def Spec.topMap {R S : CommRingCat.{u}} (f : R ⟶ S) : Spec.topObj S ⟶ Spec.topObj R := - PrimeSpectrum.comap f + PrimeSpectrum.comap f.hom @[simp] theorem Spec.topMap_id (R : CommRingCat.{u}) : Spec.topMap (𝟙 R) = 𝟙 (Spec.topObj R) := @@ -95,9 +95,9 @@ def Spec.sheafedSpaceMap {R S : CommRingCat.{u}} (f : R ⟶ S) : Spec.sheafedSpaceObj S ⟶ Spec.sheafedSpaceObj R where base := Spec.topMap f c := - { app := fun U => - comap f (unop U) ((TopologicalSpace.Opens.map (Spec.topMap f)).obj (unop U)) fun _ => id - naturality := fun {_ _} _ => RingHom.ext fun _ => Subtype.eq <| funext fun _ => rfl } + { app := fun U => CommRingCat.ofHom <| + comap f.hom (unop U) ((TopologicalSpace.Opens.map (Spec.topMap f)).obj (unop U)) fun _ => id + naturality := fun {_ _} _ => by ext; rfl } @[simp] theorem Spec.sheafedSpaceMap_id {R : CommRingCat.{u}} : @@ -105,8 +105,9 @@ theorem Spec.sheafedSpaceMap_id {R : CommRingCat.{u}} : AlgebraicGeometry.PresheafedSpace.Hom.ext _ _ (Spec.topMap_id R) <| by ext dsimp - erw [comap_id (by simp)] + rw [comap_id (by simp)] simp + rfl theorem Spec.sheafedSpaceMap_comp {R S T : CommRingCat.{u}} (f : R ⟶ S) (g : S ⟶ T) : Spec.sheafedSpaceMap (f ≫ g) = Spec.sheafedSpaceMap g ≫ Spec.sheafedSpaceMap f := @@ -194,9 +195,9 @@ lemma Spec.locallyRingedSpaceObj_presheaf_map' (R : Type u) [CommRing R] {U V} ( @[elementwise] theorem stalkMap_toStalk {R S : CommRingCat.{u}} (f : R ⟶ S) (p : PrimeSpectrum S) : - toStalk R (PrimeSpectrum.comap f p) ≫ (Spec.sheafedSpaceMap f).stalkMap p = + toStalk R (PrimeSpectrum.comap f.hom p) ≫ (Spec.sheafedSpaceMap f).stalkMap p = f ≫ toStalk S p := by - rw [← toOpen_germ S ⊤ p trivial, ← toOpen_germ R ⊤ (PrimeSpectrum.comap f p) trivial, + rw [← toOpen_germ S ⊤ p trivial, ← toOpen_germ R ⊤ (PrimeSpectrum.comap f.hom p) trivial, Category.assoc] erw [PresheafedSpace.stalkMap_germ (Spec.sheafedSpaceMap f) ⊤ p trivial] rw [Spec.sheafedSpaceMap_c_app] @@ -208,20 +209,27 @@ to the induced local ring homomorphism `Localization.localRingHom`. -/ @[elementwise] theorem localRingHom_comp_stalkIso {R S : CommRingCat.{u}} (f : R ⟶ S) (p : PrimeSpectrum S) : - (stalkIso R (PrimeSpectrum.comap f p)).hom ≫ - @CategoryStruct.comp _ _ - (CommRingCat.of (Localization.AtPrime (PrimeSpectrum.comap f p).asIdeal)) - (CommRingCat.of (Localization.AtPrime p.asIdeal)) _ - (Localization.localRingHom (PrimeSpectrum.comap f p).asIdeal p.asIdeal f rfl) - (stalkIso S p).inv = + (stalkIso R (PrimeSpectrum.comap f.hom p)).hom ≫ + (CommRingCat.ofHom (Localization.localRingHom (PrimeSpectrum.comap f.hom p).asIdeal p.asIdeal + f.hom rfl)) ≫ + (stalkIso S p).inv = (Spec.sheafedSpaceMap f).stalkMap p := - (stalkIso R (PrimeSpectrum.comap f p)).eq_inv_comp.mp <| - (stalkIso S p).comp_inv_eq.mpr <| - Localization.localRingHom_unique _ _ _ _ fun x => by + (stalkIso R (PrimeSpectrum.comap f.hom p)).eq_inv_comp.mp <| + (stalkIso S p).comp_inv_eq.mpr <| CommRingCat.hom_ext <| + Localization.localRingHom_unique _ _ _ (PrimeSpectrum.comap_asIdeal _ _) fun x => by -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 and https://github.com/leanprover-community/mathlib4/pull/8386 - rw [stalkIso_hom, stalkIso_inv] - erw [comp_apply, comp_apply, localizationToStalk_of, stalkMap_toStalk_apply f p x, - stalkToFiberRingHom_toStalk] + rw [stalkIso_hom, stalkIso_inv, CommRingCat.comp_apply, CommRingCat.comp_apply, + localizationToStalk_of] + erw [stalkMap_toStalk_apply f p x, stalkToFiberRingHom_toStalk] + rfl + +/-- Version of `localRingHom_comp_stalkIso_apply` using `CommRingCat.Hom.hom` -/ +theorem localRingHom_comp_stalkIso_apply' {R S : CommRingCat.{u}} (f : R ⟶ S) (p : PrimeSpectrum S) + (x) : + (stalkIso S p).inv ((Localization.localRingHom (PrimeSpectrum.comap f.hom p).asIdeal p.asIdeal + f.hom rfl) ((stalkIso R (PrimeSpectrum.comap f.hom p)).hom x)) = + (Spec.sheafedSpaceMap f).stalkMap p x := + localRingHom_comp_stalkIso_apply _ _ _ /-- The induced map of a ring homomorphism on the prime spectra, as a morphism of locally ringed spaces. @@ -237,14 +245,14 @@ def Spec.locallyRingedSpaceMap {R S : CommRingCat.{u}} (f : R ⟶ S) : #adaptation_note /-- nightly-2024-04-01 It's this `erw` that is blowing up. The implicit arguments differ significantly. -/ - erw [← localRingHom_comp_stalkIso_apply] at ha - replace ha := (isUnit_map_iff (stalkIso S p).inv _).mp ha - -- Porting note: `f` had to be made explicit + erw [← localRingHom_comp_stalkIso_apply' f p a] at ha + + have : IsLocalHom (stalkIso (↑S) p).inv.hom := isLocalHom_of_isIso _ + replace ha := (isUnit_map_iff (stalkIso S p).inv.hom _).mp ha replace ha := IsLocalHom.map_nonunit - (f := (Localization.localRingHom (PrimeSpectrum.comap f p).asIdeal p.asIdeal f _)) - ((stalkIso R ((PrimeSpectrum.comap f) p)).hom a) ha - convert RingHom.isUnit_map (stalkIso R (PrimeSpectrum.comap f p)).inv ha - erw [← comp_apply, Iso.hom_inv_id, id_apply] + ((stalkIso R ((PrimeSpectrum.comap f.hom) p)).hom a) ha + convert RingHom.isUnit_map (stalkIso R (PrimeSpectrum.comap f.hom p)).inv.hom ha + rw [← CommRingCat.comp_apply, Iso.hom_inv_id, CommRingCat.id_apply] @[simp] theorem Spec.locallyRingedSpaceMap_id (R : CommRingCat.{u}) : @@ -277,7 +285,7 @@ def toSpecΓ (R : CommRingCat.{u}) : R ⟶ Γ.obj (op (Spec.toLocallyRingedSpace StructureSheaf.toOpen R ⊤ -- These lemmas have always been bad (https://github.com/leanprover-community/mathlib4/issues/7657), but https://github.com/leanprover/lean4/pull/2644 made `simp` start noticing -attribute [nolint simpNF] AlgebraicGeometry.toSpecΓ_apply_coe +attribute [nolint simpNF] AlgebraicGeometry.toSpecΓ_hom_apply_coe instance isIso_toSpecΓ (R : CommRingCat.{u}) : IsIso (toSpecΓ R) := by cases R; apply StructureSheaf.isIso_to_global @@ -286,7 +294,8 @@ instance isIso_toSpecΓ (R : CommRingCat.{u}) : IsIso (toSpecΓ R) := by theorem Spec_Γ_naturality {R S : CommRingCat.{u}} (f : R ⟶ S) : f ≫ toSpecΓ S = toSpecΓ R ≫ Γ.map (Spec.toLocallyRingedSpace.map f.op).op := by -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): `ext` failed to pick up one of the three lemmas - refine RingHom.ext fun x => Subtype.ext <| funext fun x' => ?_; symm + ext : 2 + refine Subtype.ext <| funext fun x' => ?_; symm apply Localization.localRingHom_to_map /-- The counit (`SpecΓIdentity.inv.op`) of the adjunction `Γ ⊣ Spec` is an isomorphism. -/ @@ -335,11 +344,11 @@ theorem toPushforwardStalk_comp : exact Spec_Γ_naturality_assoc f _ instance : Algebra R ((Spec.topMap f _* (structureSheaf S).1).stalk p) := - (f ≫ StructureSheaf.toPushforwardStalk f p).toAlgebra + (f ≫ StructureSheaf.toPushforwardStalk f p).hom.toAlgebra theorem algebraMap_pushforward_stalk : algebraMap R ((Spec.topMap f _* (structureSheaf S).1).stalk p) = - f ≫ StructureSheaf.toPushforwardStalk f p := + (f ≫ StructureSheaf.toPushforwardStalk f p).hom := rfl variable (R S) @@ -351,22 +360,25 @@ algebra `R ⟶ S` and some `p : Spec R`. -/ @[simps!] def toPushforwardStalkAlgHom : - S →ₐ[R] (Spec.topMap (algebraMap R S) _* (structureSheaf S).1).stalk p := - { StructureSheaf.toPushforwardStalk (algebraMap R S) p with commutes' := fun _ => rfl } + S →ₐ[R] (Spec.topMap (CommRingCat.ofHom (algebraMap R S)) _* (structureSheaf S).1).stalk p := + { (StructureSheaf.toPushforwardStalk (CommRingCat.ofHom (algebraMap R S)) p).hom with + commutes' := fun _ => rfl } theorem isLocalizedModule_toPushforwardStalkAlgHom_aux (y) : ∃ x : S × p.asIdeal.primeCompl, x.2 • y = toPushforwardStalkAlgHom R S p x.1 := by obtain ⟨U, hp, s, e⟩ := TopCat.Presheaf.germ_exist -- Porting note: originally the first variable does not need to be explicit - (Spec.topMap (algebraMap ↑R ↑S) _* (structureSheaf S).val) _ y + (Spec.topMap (CommRingCat.ofHom (algebraMap ↑R ↑S)) _* (structureSheaf S).val) _ y obtain ⟨_, ⟨r, rfl⟩, hpr : p ∈ PrimeSpectrum.basicOpen r, hrU : PrimeSpectrum.basicOpen r ≤ U⟩ := PrimeSpectrum.isTopologicalBasis_basic_opens.exists_subset_of_mem_open (show p ∈ U from hp) U.2 change PrimeSpectrum.basicOpen r ≤ U at hrU replace e := - ((Spec.topMap (algebraMap R S) _* (structureSheaf S).1).germ_res_apply (homOfLE hrU) - p hpr _).trans e - set s' := (Spec.topMap (algebraMap R S) _* (structureSheaf S).1).map (homOfLE hrU).op s with h - replace e : ((Spec.topMap (algebraMap R S) _* (structureSheaf S).val).germ _ p hpr) s' = y := by + ((Spec.topMap (CommRingCat.ofHom (algebraMap R S)) _* (structureSheaf S).1).germ_res_apply + (homOfLE hrU) p hpr _).trans e + set s' := (Spec.topMap (CommRingCat.ofHom (algebraMap R S)) _* (structureSheaf S).1).map + (homOfLE hrU).op s with h + replace e : ((Spec.topMap (CommRingCat.ofHom (algebraMap R S)) _* (structureSheaf S).val).germ _ + p hpr) s' = y := by rw [h]; exact e clear_value s'; clear! U obtain ⟨⟨s, ⟨_, n, rfl⟩⟩, hsn⟩ := @@ -374,13 +386,14 @@ theorem isLocalizedModule_toPushforwardStalkAlgHom_aux (y) : (StructureSheaf.IsLocalization.to_basicOpen S <| algebraMap R S r) s' refine ⟨⟨s, ⟨r, hpr⟩ ^ n⟩, ?_⟩ rw [Submonoid.smul_def, Algebra.smul_def, algebraMap_pushforward_stalk, toPushforwardStalk, - comp_apply, comp_apply] + CommRingCat.comp_apply, CommRingCat.comp_apply] iterate 2 - erw [← (Spec.topMap (algebraMap R S) _* (structureSheaf S).1).germ_res_apply (homOfLE le_top) - p hpr] + erw [← (Spec.topMap (CommRingCat.ofHom (algebraMap R S)) _* (structureSheaf S).1).germ_res_apply + (homOfLE le_top) p hpr] rw [← e] -- Porting note: without this `change`, Lean doesn't know how to rewrite `map_mul` - let f := TopCat.Presheaf.germ (Spec.topMap (algebraMap R S) _* (structureSheaf S).val) _ p hpr + let f := TopCat.Presheaf.germ (Spec.topMap (CommRingCat.ofHom (algebraMap R S)) _* + (structureSheaf S).val) _ p hpr change f _ * f _ = f _ rw [← map_mul, mul_comm] dsimp only [Subtype.coe_mk] at hsn @@ -392,15 +405,16 @@ instance isLocalizedModule_toPushforwardStalkAlgHom : apply IsLocalizedModule.mkOfAlgebra · intro x hx; rw [algebraMap_pushforward_stalk, toPushforwardStalk_comp] change IsUnit ((TopCat.Presheaf.stalkFunctor CommRingCat p).map - (Spec.sheafedSpaceMap (algebraMap ↑R ↑S)).c _) + (Spec.sheafedSpaceMap (CommRingCat.ofHom (algebraMap ↑R ↑S))).c _) exact (IsLocalization.map_units ((structureSheaf R).presheaf.stalk p) ⟨x, hx⟩).map _ · apply isLocalizedModule_toPushforwardStalkAlgHom_aux · intro x hx - rw [toPushforwardStalkAlgHom_apply, ← (toPushforwardStalk (algebraMap R S) p).map_zero, + rw [toPushforwardStalkAlgHom_apply, + ← (toPushforwardStalk (CommRingCat.ofHom (algebraMap ↑R ↑S)) p).hom.map_zero, toPushforwardStalk] at hx -- Porting note: this `change` is manually rewriting `comp_apply` - change _ = (TopCat.Presheaf.germ (Spec.topMap (algebraMap ↑R ↑S) _* (structureSheaf ↑S).val) - ⊤ p trivial (toOpen S ⊤ 0)) at hx + change _ = (TopCat.Presheaf.germ (Spec.topMap (CommRingCat.ofHom (algebraMap ↑R ↑S)) _* + (structureSheaf ↑S).val) ⊤ p trivial (toOpen S ⊤ 0)) at hx rw [map_zero] at hx change (forget CommRingCat).map _ _ = (forget _).map _ _ at hx obtain ⟨U, hpU, i₁, i₂, e⟩ := TopCat.Presheaf.germ_eq _ _ _ _ _ _ hx @@ -408,7 +422,8 @@ instance isLocalizedModule_toPushforwardStalkAlgHom : PrimeSpectrum.isTopologicalBasis_basic_opens.exists_subset_of_mem_open (show p ∈ U.1 from hpU) U.2 change PrimeSpectrum.basicOpen r ≤ U at hrU - apply_fun (Spec.topMap (algebraMap R S) _* (structureSheaf S).1).map (homOfLE hrU).op at e + apply_fun (Spec.topMap (CommRingCat.ofHom (algebraMap R S)) _* (structureSheaf S).1).map + (homOfLE hrU).op at e simp only [Functor.op_map, map_zero, ← comp_apply, toOpen_res] at e have : toOpen S (PrimeSpectrum.basicOpen <| algebraMap R S r) x = 0 := by refine Eq.trans ?_ e; rfl diff --git a/Mathlib/AlgebraicGeometry/SpreadingOut.lean b/Mathlib/AlgebraicGeometry/SpreadingOut.lean index a327e1ae7ee02..c66c19a1f9172 100644 --- a/Mathlib/AlgebraicGeometry/SpreadingOut.lean +++ b/Mathlib/AlgebraicGeometry/SpreadingOut.lean @@ -66,7 +66,7 @@ lemma injective_germ_basicOpen (U : X.Opens) (hU : IsAffineOpen U) have := hU.isLocalization_basicOpen f obtain ⟨t, s, rfl⟩ := IsLocalization.mk'_surjective (.powers f) t rw [← RingHom.mem_ker, IsLocalization.mk'_eq_mul_mk'_one, Ideal.mul_unit_mem_iff_mem, - RingHom.mem_ker, RingHom.algebraMap_toAlgebra, X.presheaf.germ_res_apply] at ht + RingHom.mem_ker, RingHom.algebraMap_toAlgebra, CommRingCat.germ_res_apply] at ht swap; · exact @isUnit_of_invertible _ _ _ (@IsLocalization.invertible_mk'_one ..) rw [H _ ht, IsLocalization.mk'_zero] @@ -210,7 +210,7 @@ lemma spread_out_unique_of_isGermInjective {x : X} [X.IsGermInjectiveAt x] congr 2 apply this <;> simp rintro U V rfl rfl - have := ConcreteCategory.mono_of_injective _ HU + have := ConcreteCategory.mono_of_injective (C := CommRingCat) _ HU rw [← cancel_mono (X.presheaf.germ U x hxU)] simp only [Scheme.Hom.appLE, Category.assoc, X.presheaf.germ_res', ← Scheme.stalkMap_germ, H] simp only [TopCat.Presheaf.germ_stalkSpecializes_assoc, Scheme.stalkMap_germ] @@ -231,23 +231,24 @@ lemma spread_out_unique_of_isGermInjective' {x : X} [X.IsGermInjectiveAt x] lemma exists_lift_of_germInjective_aux {U : X.Opens} {x : X} (hxU) (φ : A ⟶ X.presheaf.stalk x) (φRA : R ⟶ A) (φRX : R ⟶ Γ(X, U)) - (hφRA : RingHom.FiniteType φRA) + (hφRA : RingHom.FiniteType φRA.hom) (e : φRA ≫ φ = φRX ≫ X.presheaf.germ U x hxU) : ∃ (V : X.Opens) (hxV : x ∈ V), - V ≤ U ∧ RingHom.range φ ≤ RingHom.range (X.presheaf.germ V x hxV) := by - letI := φRA.toAlgebra + V ≤ U ∧ RingHom.range φ.hom ≤ RingHom.range (X.presheaf.germ V x hxV).hom := by + letI := φRA.hom.toAlgebra obtain ⟨s, hs⟩ := hφRA choose W hxW f hf using fun t ↦ X.presheaf.germ_exist x (φ t) have H : x ∈ s.inf W ⊓ U := by rw [← SetLike.mem_coe, TopologicalSpace.Opens.coe_inf, TopologicalSpace.Opens.coe_finset_inf] exact ⟨by simpa using fun x _ ↦ hxW x, hxU⟩ refine ⟨s.inf W ⊓ U, H, inf_le_right, ?_⟩ - letI := φRX.toAlgebra - letI := (φRX ≫ X.presheaf.germ U x hxU).toAlgebra - letI := (φRX ≫ X.presheaf.map (homOfLE (inf_le_right (a := s.inf W))).op).toAlgebra - let φ' : A →ₐ[R] X.presheaf.stalk x := { φ with commutes' := DFunLike.congr_fun e } + letI := φRX.hom.toAlgebra + letI := (φRX ≫ X.presheaf.germ U x hxU).hom.toAlgebra + letI := (φRX ≫ X.presheaf.map (homOfLE (inf_le_right (a := s.inf W))).op).hom.toAlgebra + let φ' : A →ₐ[R] X.presheaf.stalk x := + { φ.hom with commutes' := DFunLike.congr_fun (congr_arg CommRingCat.Hom.hom e) } let ψ : Γ(X, s.inf W ⊓ U) →ₐ[R] X.presheaf.stalk x := - { X.presheaf.germ _ x H with commutes' := fun x ↦ X.presheaf.germ_res_apply _ _ _ _ } + { (X.presheaf.germ _ x H).hom with commutes' := fun x ↦ X.presheaf.germ_res_apply _ _ _ _ } show AlgHom.range φ' ≤ AlgHom.range ψ rw [← Algebra.map_top, ← hs, AlgHom.map_adjoin, Algebra.adjoin_le_iff] rintro _ ⟨i, hi, rfl : φ i = _⟩ @@ -267,19 +268,19 @@ such that `R` is of finite type over `A`, we may lift `A ⟶ 𝒪_{X, x}` to som -/ lemma exists_lift_of_germInjective {x : X} [X.IsGermInjectiveAt x] {U : X.Opens} (hxU : x ∈ U) (φ : A ⟶ X.presheaf.stalk x) (φRA : R ⟶ A) (φRX : R ⟶ Γ(X, U)) - (hφRA : RingHom.FiniteType φRA) + (hφRA : RingHom.FiniteType φRA.hom) (e : φRA ≫ φ = φRX ≫ X.presheaf.germ U x hxU) : ∃ (V : X.Opens) (hxV : x ∈ V) (φ' : A ⟶ Γ(X, V)) (i : V ≤ U), IsAffineOpen V ∧ φ = φ' ≫ X.presheaf.germ V x hxV ∧ φRX ≫ X.presheaf.map i.hom.op = φRA ≫ φ' := by obtain ⟨V, hxV, iVU, hV⟩ := exists_lift_of_germInjective_aux hxU φ φRA φRX hφRA e obtain ⟨V', hxV', hV', iV'V, H⟩ := X.exists_le_and_germ_injective x V hxV let f := X.presheaf.germ V' x hxV' - have hf' : RingHom.range (X.presheaf.germ V x hxV) ≤ RingHom.range f := by + have hf' : RingHom.range (X.presheaf.germ V x hxV).hom ≤ RingHom.range f.hom := by rw [← X.presheaf.germ_res iV'V.hom _ hxV'] exact Set.range_comp_subset_range (X.presheaf.map iV'V.hom.op) f let e := RingEquiv.ofLeftInverse H.hasLeftInverse.choose_spec refine ⟨V', hxV', CommRingCat.ofHom (e.symm.toRingHom.comp - (φ.codRestrict _ (fun x ↦ hf' (hV ⟨x, rfl⟩)))), iV'V.trans iVU, hV', ?_, ?_⟩ + (φ.hom.codRestrict _ (fun x ↦ hf' (hV ⟨x, rfl⟩)))), iV'V.trans iVU, hV', ?_, ?_⟩ · ext a show φ a = (e (e.symm _)).1 simp only [RingEquiv.apply_symm_apply] @@ -290,7 +291,7 @@ lemma exists_lift_of_germInjective {x : X} [X.IsGermInjectiveAt x] {U : X.Opens} rw [RingEquiv.apply_symm_apply] ext show X.presheaf.germ _ _ _ (X.presheaf.map _ _) = (φRA ≫ φ) a - rw [X.presheaf.germ_res_apply, ‹φRA ≫ φ = _›] + rw [CommRingCat.germ_res_apply, ‹φRA ≫ φ = _›] rfl /-- diff --git a/Mathlib/AlgebraicGeometry/Stalk.lean b/Mathlib/AlgebraicGeometry/Stalk.lean index e06eb1b7ef0ca..54411c0594348 100644 --- a/Mathlib/AlgebraicGeometry/Stalk.lean +++ b/Mathlib/AlgebraicGeometry/Stalk.lean @@ -79,7 +79,7 @@ instance IsAffineOpen.fromSpecStalk_isPreimmersion {X : Scheme.{u}} {U : Opens X (hU : IsAffineOpen U) (x : X) (hx : x ∈ U) : IsPreimmersion (hU.fromSpecStalk hx) := by dsimp [IsAffineOpen.fromSpecStalk] haveI : IsPreimmersion (Spec.map (X.presheaf.germ U x hx)) := - letI : Algebra Γ(X, U) (X.presheaf.stalk x) := (X.presheaf.germ U x hx).toAlgebra + letI : Algebra Γ(X, U) (X.presheaf.stalk x) := (X.presheaf.germ U x hx).hom.toAlgebra haveI := hU.isLocalization_stalk ⟨x, hx⟩ IsPreimmersion.of_isLocalization (R := Γ(X, U)) (S := X.presheaf.stalk x) (hU.primeIdealOf ⟨x, hx⟩).asIdeal.primeCompl @@ -262,8 +262,17 @@ def stalkClosedPointTo : f.stalkMap (closedPoint R) ≫ (stalkClosedPointIso R).hom instance isLocalHom_stalkClosedPointTo : - IsLocalHom (stalkClosedPointTo f) := - inferInstanceAs <| IsLocalHom (f.stalkMap (closedPoint R) ≫ (stalkClosedPointIso R).hom) + IsLocalHom (stalkClosedPointTo f).hom := + inferInstanceAs <| IsLocalHom (f.stalkMap (closedPoint R) ≫ (stalkClosedPointIso R).hom).hom + +/-- Copy of `isLocalHom_stalkClosedPointTo` which unbundles the comm ring. + +Useful for use in combination with `CommRingCat.of K` for a field `K`. +-/ +instance isLocalHom_stalkClosedPointTo' {R : Type u} [CommRing R] [IsLocalRing R] + (f : Spec (.of R) ⟶ X) : + IsLocalHom (stalkClosedPointTo f).hom := + isLocalHom_stalkClosedPointTo f lemma preimage_eq_top_of_closedPoint_mem {U : Opens X} (hU : f.base (closedPoint R) ∈ U) : f ⁻¹ᵁ U = ⊤ := @@ -295,7 +304,7 @@ lemma germ_stalkClosedPointTo (U : Opens X) (hU : f.base (closedPoint R) ∈ U) @[reassoc] lemma germ_stalkClosedPointTo_Spec_fromSpecStalk - {x : X} (f : X.presheaf.stalk x ⟶ R) [IsLocalHom f] (U : Opens X) (hU) : + {x : X} (f : X.presheaf.stalk x ⟶ R) [IsLocalHom f.hom] (U : Opens X) (hU) : X.presheaf.germ U _ hU ≫ stalkClosedPointTo (Spec.map f ≫ X.fromSpecStalk x) = X.presheaf.germ U x (by simpa using hU) ≫ f := by have : (Spec.map f ≫ X.fromSpecStalk x).base (closedPoint R) = x := by @@ -339,7 +348,7 @@ variable {R} omit [IsLocalRing R] in /-- useful lemma for applications of `SpecToEquivOfLocalRing` -/ lemma SpecToEquivOfLocalRing_eq_iff - {f₁ f₂ : Σ x, { f : X.presheaf.stalk x ⟶ R // IsLocalHom f }} : + {f₁ f₂ : Σ x, { f : X.presheaf.stalk x ⟶ R // IsLocalHom f.hom }} : f₁ = f₂ ↔ ∃ h₁ : f₁.1 = f₂.1, f₁.2.1 = (X.presheaf.stalkCongr (by rw [h₁]; rfl)).hom ≫ f₂.2.1 := by constructor @@ -358,7 +367,7 @@ Given a local ring `R` and scheme `X`, morphisms `Spec R ⟶ X` corresponds to p @[simps] noncomputable def SpecToEquivOfLocalRing : - (Spec R ⟶ X) ≃ Σ x, { f : X.presheaf.stalk x ⟶ R // IsLocalHom f } where + (Spec R ⟶ X) ≃ Σ x, { f : X.presheaf.stalk x ⟶ R // IsLocalHom f.hom } where toFun f := ⟨f.base (closedPoint R), Scheme.stalkClosedPointTo f, inferInstance⟩ invFun xf := Spec.map xf.2.1 ≫ X.fromSpecStalk xf.1 left_inv := Scheme.Spec_stalkClosedPointTo_fromSpecStalk diff --git a/Mathlib/AlgebraicGeometry/StructureSheaf.lean b/Mathlib/AlgebraicGeometry/StructureSheaf.lean index 7a8498fdd5845..8d35ebe6a7520 100644 --- a/Mathlib/AlgebraicGeometry/StructureSheaf.lean +++ b/Mathlib/AlgebraicGeometry/StructureSheaf.lean @@ -226,7 +226,7 @@ structure presheaf. @[simps] def structurePresheafInCommRing : Presheaf CommRingCat (PrimeSpectrum.Top R) where obj U := CommRingCat.of ((structureSheafInType R).1.obj U) - map {_ _} i := + map {_ _} i := CommRingCat.ofHom { toFun := (structureSheafInType R).1.map i map_zero' := rfl map_add' := fun _ _ => rfl @@ -234,7 +234,7 @@ def structurePresheafInCommRing : Presheaf CommRingCat (PrimeSpectrum.Top R) whe map_mul' := fun _ _ => rfl } -- These lemmas have always been bad (https://github.com/leanprover-community/mathlib4/issues/7657), but https://github.com/leanprover/lean4/pull/2644 made `simp` start noticing -attribute [nolint simpNF] AlgebraicGeometry.structurePresheafInCommRing_map_apply +attribute [nolint simpNF] AlgebraicGeometry.structurePresheafInCommRing_map_hom_apply /-- Some glue, verifying that the structure presheaf valued in `CommRingCat` agrees with the `Type` valued structure presheaf. @@ -377,15 +377,15 @@ theorem const_mul_cancel' (f g₁ g₂ : R) (U hu₁ hu₂) : /-- The canonical ring homomorphism interpreting an element of `R` as a section of the structure sheaf. -/ def toOpen (U : Opens (PrimeSpectrum.Top R)) : - CommRingCat.of R ⟶ (structureSheaf R).1.obj (op U) where - toFun f := - ⟨fun _ => algebraMap R _ f, fun x => - ⟨U, x.2, 𝟙 _, f, 1, fun y => - ⟨(Ideal.ne_top_iff_one _).1 y.1.2.1, by rw [RingHom.map_one, mul_one]⟩⟩⟩ - map_one' := Subtype.eq <| funext fun _ => RingHom.map_one _ - map_mul' _ _ := Subtype.eq <| funext fun _ => RingHom.map_mul _ _ _ - map_zero' := Subtype.eq <| funext fun _ => RingHom.map_zero _ - map_add' _ _ := Subtype.eq <| funext fun _ => RingHom.map_add _ _ _ + CommRingCat.of R ⟶ (structureSheaf R).1.obj (op U) := CommRingCat.ofHom + { toFun f := + ⟨fun _ => algebraMap R _ f, fun x => + ⟨U, x.2, 𝟙 _, f, 1, fun y => + ⟨(Ideal.ne_top_iff_one _).1 y.1.2.1, by rw [RingHom.map_one, mul_one]⟩⟩⟩ + map_one' := Subtype.eq <| funext fun _ => RingHom.map_one _ + map_mul' _ _ := Subtype.eq <| funext fun _ => RingHom.map_mul _ _ _ + map_zero' := Subtype.eq <| funext fun _ => RingHom.map_zero _ + map_add' _ _ := Subtype.eq <| funext fun _ => RingHom.map_add _ _ _ } @[simp] theorem toOpen_res (U V : Opens (PrimeSpectrum.Top R)) (i : V ⟶ U) : @@ -436,7 +436,8 @@ theorem isUnit_toStalk (x : PrimeSpectrum.Top R) (f : x.asIdeal.primeCompl) : of the structure sheaf at the point `p`. -/ def localizationToStalk (x : PrimeSpectrum.Top R) : CommRingCat.of (Localization.AtPrime x.asIdeal) ⟶ (structureSheaf R).presheaf.stalk x := - show Localization.AtPrime x.asIdeal →+* _ from IsLocalization.lift (isUnit_toStalk R x) + CommRingCat.ofHom <| + show Localization.AtPrime x.asIdeal →+* _ from IsLocalization.lift (isUnit_toStalk R x) @[simp] theorem localizationToStalk_of (x : PrimeSpectrum.Top R) (f : R) : @@ -457,12 +458,13 @@ theorem localizationToStalk_mk' (x : PrimeSpectrum.Top R) (f : R) (s : x.asIdeal implemented as a subtype of dependent functions to localizations at prime ideals, and evaluates the section on the point corresponding to a given prime ideal. -/ def openToLocalization (U : Opens (PrimeSpectrum.Top R)) (x : PrimeSpectrum.Top R) (hx : x ∈ U) : - (structureSheaf R).1.obj (op U) ⟶ CommRingCat.of (Localization.AtPrime x.asIdeal) where - toFun s := (s.1 ⟨x, hx⟩ : _) - map_one' := rfl - map_mul' _ _ := rfl - map_zero' := rfl - map_add' _ _ := rfl + (structureSheaf R).1.obj (op U) ⟶ CommRingCat.of (Localization.AtPrime x.asIdeal) := + CommRingCat.ofHom + { toFun s := (s.1 ⟨x, hx⟩ : _) + map_one' := rfl + map_mul' _ _ := rfl + map_zero' := rfl + map_add' _ _ := rfl } @[simp] theorem coe_openToLocalization (U : Opens (PrimeSpectrum.Top R)) (x : PrimeSpectrum.Top R) @@ -498,21 +500,19 @@ theorem germ_comp_stalkToFiberRingHom theorem stalkToFiberRingHom_germ (U : Opens (PrimeSpectrum.Top R)) (x : PrimeSpectrum.Top R) (hx : x ∈ U) (s : (structureSheaf R).1.obj (op U)) : stalkToFiberRingHom R x ((structureSheaf R).presheaf.germ U x hx s) = s.1 ⟨x, hx⟩ := - RingHom.ext_iff.mp (germ_comp_stalkToFiberRingHom R U x hx) s + RingHom.ext_iff.mp (CommRingCat.hom_ext_iff.mp (germ_comp_stalkToFiberRingHom R U x hx)) s @[deprecated (since := "2024-07-30")] alias stalkToFiberRingHom_germ' := stalkToFiberRingHom_germ @[simp] theorem toStalk_comp_stalkToFiberRingHom (x : PrimeSpectrum.Top R) : - -- Porting note: now `algebraMap _ _` needs to be explicitly typed - toStalk R x ≫ stalkToFiberRingHom R x = algebraMap R (Localization.AtPrime x.asIdeal) := by + toStalk R x ≫ stalkToFiberRingHom R x = CommRingCat.ofHom (algebraMap _ _) := by rw [toStalk, Category.assoc, germ_comp_stalkToFiberRingHom]; rfl @[simp] theorem stalkToFiberRingHom_toStalk (x : PrimeSpectrum.Top R) (f : R) : - -- Porting note: now `algebraMap _ _` needs to be explicitly typed - stalkToFiberRingHom R x (toStalk R x f) = algebraMap R (Localization.AtPrime x.asIdeal) f := - RingHom.ext_iff.1 (toStalk_comp_stalkToFiberRingHom R x) _ + stalkToFiberRingHom R x (toStalk R x f) = algebraMap _ _ f := + RingHom.ext_iff.1 (CommRingCat.hom_ext_iff.mp (toStalk_comp_stalkToFiberRingHom R x)) _ /-- The ring isomorphism between the stalk of the structure sheaf of `R` at a point `p` corresponding to a prime ideal in `R` and the localization of `R` at `p`. -/ @@ -522,19 +522,24 @@ def stalkIso (x : PrimeSpectrum.Top R) : hom := stalkToFiberRingHom R x inv := localizationToStalk R x hom_inv_id := by - ext U hxU s - dsimp only [Category.comp_id, CommRingCat.forget_obj, - CommRingCat.coe_of, CommRingCat.comp_apply] + apply stalk_hom_ext + intro U hxU + ext s + dsimp only [CommRingCat.hom_comp, RingHom.coe_comp, Function.comp_apply, CommRingCat.hom_id, + RingHom.coe_id, id_eq] rw [stalkToFiberRingHom_germ] obtain ⟨V, hxV, iVU, f, g, (hg : V ≤ PrimeSpectrum.basicOpen _), hs⟩ := exists_const _ _ s x hxU rw [← res_apply R U V iVU s ⟨x, hxV⟩, ← hs, const_apply, localizationToStalk_mk'] refine (structureSheaf R).presheaf.germ_ext V hxV (homOfLE hg) iVU ?_ + -- Replace the `ConcreteCategory.instFunLike` instance with `CommRingCat.hom`: + show (structureSheaf R).presheaf.map (homOfLE hg).op _ = + (structureSheaf R).presheaf.map iVU.op s rw [← hs, res_const'] - inv_hom_id := + inv_hom_id := CommRingCat.hom_ext <| @IsLocalization.ringHom_ext R _ x.asIdeal.primeCompl (Localization.AtPrime x.asIdeal) _ _ (Localization.AtPrime x.asIdeal) _ _ - (RingHom.comp (stalkToFiberRingHom R x) (localizationToStalk R x)) + (RingHom.comp (stalkToFiberRingHom R x).hom (localizationToStalk R x).hom) (RingHom.id (Localization.AtPrime _)) <| by ext f rw [RingHom.comp_apply, RingHom.comp_apply, localizationToStalk_of, @@ -543,13 +548,13 @@ def stalkIso (x : PrimeSpectrum.Top R) : instance (x : PrimeSpectrum R) : IsIso (stalkToFiberRingHom R x) := (stalkIso R x).isIso_hom -instance (x : PrimeSpectrum R) : IsLocalHom (stalkToFiberRingHom R x) := +instance (x : PrimeSpectrum R) : IsLocalHom (stalkToFiberRingHom R x).hom := isLocalHom_of_isIso _ instance (x : PrimeSpectrum R) : IsIso (localizationToStalk R x) := (stalkIso R x).isIso_inv -instance (x : PrimeSpectrum R) : IsLocalHom (localizationToStalk R x) := +instance (x : PrimeSpectrum R) : IsLocalHom (localizationToStalk R x).hom := isLocalHom_of_isIso _ @[simp, reassoc] @@ -578,7 +583,7 @@ theorem toBasicOpen_mk' (s f : R) (g : Submonoid.powers s) : @[simp] theorem localization_toBasicOpen (f : R) : RingHom.comp (toBasicOpen R f) (algebraMap R (Localization.Away f)) = - toOpen R (PrimeSpectrum.basicOpen f) := + (toOpen R (PrimeSpectrum.basicOpen f)).hom := RingHom.ext fun g => by rw [toBasicOpen, IsLocalization.Away.lift, RingHom.comp_apply, IsLocalization.lift_eq] @@ -830,7 +835,7 @@ theorem toBasicOpen_surjective (f : R) : Function.Surjective (toBasicOpen R f) : exact ⟨⟨i, i_mem⟩, x_mem⟩ rintro ⟨i, hi⟩ dsimp - change (structureSheaf R).1.map _ _ = (structureSheaf R).1.map _ _ + change (structureSheaf R).1.map (iDh i).op _ = (structureSheaf R).1.map (iDh i).op _ rw [s_eq i hi, res_const] -- Again, `res_const` spits out an additional goal swap @@ -847,9 +852,8 @@ theorem toBasicOpen_surjective (f : R) : Function.Surjective (toBasicOpen R f) : ring instance isIso_toBasicOpen (f : R) : - IsIso (show CommRingCat.of (Localization.Away f) ⟶ _ from toBasicOpen R f) := - haveI : IsIso ((forget CommRingCat).map - (show CommRingCat.of (Localization.Away f) ⟶ _ from toBasicOpen R f)) := + IsIso (CommRingCat.ofHom (toBasicOpen R f)) := + haveI : IsIso ((forget CommRingCat).map (CommRingCat.ofHom (toBasicOpen R f))) := (isIso_iff_bijective _).mpr ⟨toBasicOpen_injective R f, toBasicOpen_surjective R f⟩ isIso_of_reflects_iso _ (forget CommRingCat) @@ -858,10 +862,10 @@ at the submonoid of powers of `f`. -/ def basicOpenIso (f : R) : (structureSheaf R).1.obj (op (PrimeSpectrum.basicOpen f)) ≅ CommRingCat.of (Localization.Away f) := - (asIso (show CommRingCat.of (Localization.Away f) ⟶ _ from toBasicOpen R f)).symm + (asIso (CommRingCat.ofHom (toBasicOpen R f))).symm instance stalkAlgebra (p : PrimeSpectrum R) : Algebra R ((structureSheaf R).presheaf.stalk p) := - (toStalk R p).toAlgebra + (toStalk R p).hom.toAlgebra @[simp] theorem stalkAlgebra_map (p : PrimeSpectrum R) (r : R) : @@ -877,13 +881,13 @@ instance IsLocalization.to_stalk (p : PrimeSpectrum R) : apply Algebra.algebra_ext intro rw [stalkAlgebra_map] - congr 1 + congr 2 change toStalk R p = _ ≫ (stalkIso R p).inv rw [Iso.eq_comp_inv] exact toStalk_comp_stalkToFiberRingHom R p instance openAlgebra (U : (Opens (PrimeSpectrum R))ᵒᵖ) : Algebra R ((structureSheaf R).val.obj U) := - (toOpen R (unop U)).toAlgebra + (toOpen R (unop U)).hom.toAlgebra @[simp] theorem openAlgebra_map (U : (Opens (PrimeSpectrum R))ᵒᵖ) (r : R) : @@ -902,20 +906,20 @@ instance IsLocalization.to_basicOpen (r : R) : exact (localization_toBasicOpen R r).symm instance to_basicOpen_epi (r : R) : Epi (toOpen R (PrimeSpectrum.basicOpen r)) := - ⟨fun _ _ h => IsLocalization.ringHom_ext (Submonoid.powers r) h⟩ + ⟨fun _ _ h => CommRingCat.hom_ext (IsLocalization.ringHom_ext (Submonoid.powers r) + (CommRingCat.hom_ext_iff.mp h))⟩ @[elementwise] theorem to_global_factors : toOpen R ⊤ = CommRingCat.ofHom (algebraMap R (Localization.Away (1 : R))) ≫ - toBasicOpen R (1 : R) ≫ + CommRingCat.ofHom (toBasicOpen R (1 : R)) ≫ (structureSheaf R).1.map (eqToHom PrimeSpectrum.basicOpen_one.symm).op := by rw [← Category.assoc] change toOpen R ⊤ = (CommRingCat.ofHom <| (toBasicOpen R 1).comp (algebraMap R (Localization.Away 1))) ≫ (structureSheaf R).1.map (eqToHom _).op - unfold CommRingCat.ofHom - rw [localization_toBasicOpen R, toOpen_res] + rw [localization_toBasicOpen R, CommRingCat.ofHom_hom, toOpen_res] instance isIso_to_global : IsIso (toOpen R ⊤) := by let hom := CommRingCat.ofHom (algebraMap R (Localization.Away (1 : R))) @@ -931,41 +935,37 @@ def globalSectionsIso : CommRingCat.of R ≅ (structureSheaf R).1.obj (op ⊤) : asIso (toOpen R ⊤) -- These lemmas have always been bad (https://github.com/leanprover-community/mathlib4/issues/7657), but https://github.com/leanprover/lean4/pull/2644 made `simp` start noticing -attribute [nolint simpNF] AlgebraicGeometry.StructureSheaf.globalSectionsIso_hom_apply_coe +attribute [nolint simpNF] AlgebraicGeometry.StructureSheaf.globalSectionsIso_hom_hom_apply_coe @[simp] theorem globalSectionsIso_hom (R : CommRingCat) : (globalSectionsIso R).hom = toOpen R ⊤ := rfl -@[simp, reassoc, elementwise] +@[simp, reassoc, elementwise nosimp] theorem toStalk_stalkSpecializes {R : Type*} [CommRing R] {x y : PrimeSpectrum R} (h : x ⤳ y) : toStalk R y ≫ (structureSheaf R).presheaf.stalkSpecializes h = toStalk R x := by dsimp [toStalk]; simp [-toOpen_germ] -@[simp, reassoc, elementwise] +@[simp, reassoc, elementwise nosimp] theorem localizationToStalk_stalkSpecializes {R : Type*} [CommRing R] {x y : PrimeSpectrum R} (h : x ⤳ y) : StructureSheaf.localizationToStalk R y ≫ (structureSheaf R).presheaf.stalkSpecializes h = CommRingCat.ofHom (PrimeSpectrum.localizationMapOfSpecializes h) ≫ StructureSheaf.localizationToStalk R x := by + ext : 1 apply IsLocalization.ringHom_ext (S := Localization.AtPrime y.asIdeal) y.asIdeal.primeCompl erw [RingHom.comp_assoc] conv_rhs => erw [RingHom.comp_assoc] dsimp [CommRingCat.ofHom, localizationToStalk, PrimeSpectrum.localizationMapOfSpecializes] rw [IsLocalization.lift_comp, IsLocalization.lift_comp, IsLocalization.lift_comp] - exact toStalk_stalkSpecializes h + exact CommRingCat.hom_ext_iff.mp (toStalk_stalkSpecializes h) -@[simp, reassoc, elementwise] +@[simp, reassoc, elementwise nosimp] theorem stalkSpecializes_stalk_to_fiber {R : Type*} [CommRing R] {x y : PrimeSpectrum R} (h : x ⤳ y) : (structureSheaf R).presheaf.stalkSpecializes h ≫ StructureSheaf.stalkToFiberRingHom R x = StructureSheaf.stalkToFiberRingHom R y ≫ - -- Porting note: `PrimeSpectrum.localizationMapOfSpecializes h` by itself is interpreted as a - -- ring homomorphism, so it is changed in a way to force it being interpreted as categorical - -- arrow. - (show CommRingCat.of (Localization.AtPrime y.asIdeal) ⟶ - CommRingCat.of (Localization.AtPrime x.asIdeal) - from PrimeSpectrum.localizationMapOfSpecializes h) := by + (CommRingCat.ofHom (PrimeSpectrum.localizationMapOfSpecializes h)) := by change _ ≫ (StructureSheaf.stalkIso R x).hom = (StructureSheaf.stalkIso R y).hom ≫ _ rw [← Iso.eq_comp_inv, Category.assoc, ← Iso.inv_comp_eq] exact localizationToStalk_stalkSpecializes h @@ -1072,7 +1072,7 @@ to OO_X(U) is the identity. -/ theorem comap_id_eq_map (U V : Opens (PrimeSpectrum.Top R)) (iVU : V ⟶ U) : (comap (RingHom.id R) U V fun _ hpV => leOfHom iVU <| hpV) = - (structureSheaf R).1.map iVU.op := + ((structureSheaf R).1.map iVU.op).hom := RingHom.ext fun s => Subtype.eq <| funext fun p => by rw [comap_apply] -- Unfortunately, we cannot use `Localization.localRingHom_id` here, because @@ -1095,7 +1095,7 @@ are not definitionally equal. -/ theorem comap_id {U V : Opens (PrimeSpectrum.Top R)} (hUV : U = V) : (comap (RingHom.id R) U V fun p hpV => by rwa [hUV, PrimeSpectrum.comap_id]) = - eqToHom (show (structureSheaf R).1.obj (op U) = _ by rw [hUV]) := by + (eqToHom (show (structureSheaf R).1.obj (op U) = _ by rw [hUV])).hom := by rw [comap_id_eq_map U V (eqToHom hUV.symm), eqToHom_op, eqToHom_map] @[simp] @@ -1119,16 +1119,19 @@ theorem comap_comp (f : R →+* S) (g : S →+* P) (U : Opens (PrimeSpectrum.Top @[elementwise, reassoc] theorem toOpen_comp_comap (f : R →+* S) (U : Opens (PrimeSpectrum.Top R)) : - (toOpen R U ≫ comap f U (Opens.comap (PrimeSpectrum.comap f) U) fun _ => id) = + (toOpen R U ≫ CommRingCat.ofHom (comap f U (Opens.comap (PrimeSpectrum.comap f) U) + fun _ => id)) = CommRingCat.ofHom f ≫ toOpen S _ := - RingHom.ext fun _ => Subtype.eq <| funext fun _ => Localization.localRingHom_to_map _ _ _ _ _ + CommRingCat.hom_ext <| RingHom.ext fun _ => Subtype.eq <| funext fun _ => + Localization.localRingHom_to_map _ _ _ _ _ lemma comap_basicOpen (f : R →+* S) (x : R) : comap f (PrimeSpectrum.basicOpen x) (PrimeSpectrum.basicOpen (f x)) (PrimeSpectrum.comap_basicOpen f x).le = IsLocalization.map (M := .powers x) (T := .powers (f x)) _ f (Submonoid.powers_le.mpr (Submonoid.mem_powers _)) := - IsLocalization.ringHom_ext (.powers x) <| by simpa using toOpen_comp_comap f _ + IsLocalization.ringHom_ext (.powers x) <| by + simpa [CommRingCat.hom_ext_iff] using toOpen_comp_comap f _ end Comap diff --git a/Mathlib/AlgebraicGeometry/ValuativeCriterion.lean b/Mathlib/AlgebraicGeometry/ValuativeCriterion.lean index 10d725b7f0215..46b7c013d0ec1 100644 --- a/Mathlib/AlgebraicGeometry/ValuativeCriterion.lean +++ b/Mathlib/AlgebraicGeometry/ValuativeCriterion.lean @@ -112,9 +112,9 @@ lemma specializingMap (H : ValuativeCriterion.Existence f) : intro x' y h let stalk_y_to_residue_x' : Y.presheaf.stalk y ⟶ X.residueField x' := Y.presheaf.stalkSpecializes h ≫ f.stalkMap x' ≫ X.residue x' - obtain ⟨A, hA, hA_local⟩ := exists_factor_valuationRing stalk_y_to_residue_x' + obtain ⟨A, hA, hA_local⟩ := exists_factor_valuationRing stalk_y_to_residue_x'.hom let stalk_y_to_A : Y.presheaf.stalk y ⟶ .of A := - CommRingCat.ofHom (stalk_y_to_residue_x'.codRestrict _ hA) + CommRingCat.ofHom (stalk_y_to_residue_x'.hom.codRestrict _ hA) have w : X.fromSpecResidueField x' ≫ f = Spec.map (CommRingCat.ofHom (algebraMap A (X.residueField x'))) ≫ Spec.map stalk_y_to_A ≫ Y.fromSpecStalk y := by @@ -130,9 +130,12 @@ lemma specializingMap (H : ValuativeCriterion.Existence f) : · rw [← Scheme.comp_base_apply, hl₂] simp only [Scheme.comp_coeBase, TopCat.coe_comp, Function.comp_apply] have : (Spec.map stalk_y_to_A).base (closedPoint A) = closedPoint (Y.presheaf.stalk y) := - comap_closedPoint (S := A) (stalk_y_to_residue_x'.codRestrict A.toSubring hA) + comap_closedPoint (S := A) (stalk_y_to_residue_x'.hom.codRestrict A.toSubring hA) rw [this, Y.fromSpecStalk_closedPoint] +instance {R S : CommRingCat} (e : R ≅ S) : IsLocalHom e.hom.hom := + isLocalHom_of_isIso _ + lemma of_specializingMap (H : (topologically @SpecializingMap).universally f) : ValuativeCriterion.Existence f := by rintro ⟨R, K, i₁, i₂, ⟨w⟩⟩ @@ -146,7 +149,9 @@ lemma of_specializingMap (H : (topologically @SpecializingMap).universally f) : (stalkClosedPointIso (.of R)).symm ≪≫ (Spec (.of R)).presheaf.stalkCongr (.of_eq h₂.symm) let α := e.hom ≫ (pullback.fst i₂ f).stalkMap x - have : IsLocalHom α := inferInstance + have : IsLocalHom e.hom.hom := isLocalHom_of_isIso e.hom + have : IsLocalHom α.hom := inferInstanceAs + (IsLocalHom (((pullback.fst i₂ f).stalkMap x).hom.comp e.hom.hom)) let β := (pullback i₂ f).presheaf.stalkSpecializes h₁ ≫ Scheme.stalkClosedPointTo lft have hαβ : α ≫ β = CommRingCat.ofHom (algebraMap R K) := by simp only [CommRingCat.coe_of, Iso.trans_hom, Iso.symm_hom, TopCat.Presheaf.stalkCongr_hom, @@ -158,19 +163,23 @@ lemma of_specializingMap (H : (topologically @SpecializingMap).universally f) : ← Scheme.comp_app_assoc lft (pullback.fst i₂ f)] rw [pullback.lift_fst] simp - have hbij := (bijective_rangeRestrict_comp_of_valuationRing (R := R) (K := K) α β hαβ) - let φ : (pullback i₂ f).presheaf.stalk x ⟶ CommRingCat.of R := - (RingEquiv.ofBijective _ hbij).symm.toRingHom.comp β.rangeRestrict + have hbij := (bijective_rangeRestrict_comp_of_valuationRing (R := R) (K := K) α.hom β.hom + (CommRingCat.hom_ext_iff.mp hαβ)) + let φ : (pullback i₂ f).presheaf.stalk x ⟶ CommRingCat.of R := CommRingCat.ofHom <| + (RingEquiv.ofBijective _ hbij).symm.toRingHom.comp β.hom.rangeRestrict have hαφ : α ≫ φ = 𝟙 _ := by ext x; exact (RingEquiv.ofBijective _ hbij).symm_apply_apply x have hαφ' : (pullback.fst i₂ f).stalkMap x ≫ φ = e.inv := by rw [← cancel_epi e.hom, ← Category.assoc, hαφ, e.hom_inv_id] have hφβ : φ ≫ CommRingCat.ofHom (algebraMap R K) = β := - hαβ ▸ RingHom.ext fun x ↦ congr_arg Subtype.val - ((RingEquiv.ofBijective _ hbij).apply_symm_apply (β.rangeRestrict x)) + hαβ ▸ CommRingCat.hom_ext (RingHom.ext fun x ↦ congr_arg Subtype.val + ((RingEquiv.ofBijective _ hbij).apply_symm_apply (β.hom.rangeRestrict x))) refine ⟨⟨⟨Spec.map ((pullback.snd i₂ f).stalkMap x ≫ φ) ≫ X.fromSpecStalk _, ?_, ?_⟩⟩⟩ · simp only [← Spec.map_comp_assoc, Category.assoc, hφβ] - simp [β, Scheme.Spec_map_stalkSpecializes_fromSpecStalk_assoc, -CommRingCat.coe_of, - Scheme.Spec_stalkClosedPointTo_fromSpecStalk_assoc, lft] + simp only [Spec.map_comp, Category.assoc, Scheme.Spec_map_stalkMap_fromSpecStalk, + Scheme.Spec_map_stalkSpecializes_fromSpecStalk_assoc, β] + -- This next line only fires as `rw`, not `simp`: + rw [Scheme.Spec_stalkClosedPointTo_fromSpecStalk_assoc] + simp [lft] · simp only [Spec.map_comp, Category.assoc, Scheme.Spec_map_stalkMap_fromSpecStalk, ← pullback.condition] rw [← Scheme.Spec_map_stalkMap_fromSpecStalk_assoc, ← Spec.map_comp_assoc, hαφ'] diff --git a/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean b/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean index d794651ff25d8..1489ae3daa633 100644 --- a/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean +++ b/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean @@ -303,10 +303,10 @@ def smoothSheafCommRing.forgetStalk (x : TopCat.of M) : (DFunLike.coe (α := ((forget CommRingCat).obj ((smoothSheafCommRing IM I M R).presheaf.obj (op ((OpenNhds.inclusion x).obj U.unop))))) - (colimit.ι ((OpenNhds.inclusion x).op ⋙ (smoothSheafCommRing IM I M R).presheaf) U)) + (colimit.ι ((OpenNhds.inclusion x).op ⋙ (smoothSheafCommRing IM I M R).presheaf) U).hom) (forgetStalk IM I M R x).hom = colimit.ι ((OpenNhds.inclusion x).op ⋙ (smoothSheaf IM I M R).presheaf) U := - ι_preservesColimitIso_hom _ _ _ + ι_preservesColimitIso_hom (forget CommRingCat) _ _ @[simp, reassoc, elementwise] lemma smoothSheafCommRing.ι_forgetStalk_inv (x : TopCat.of M) (U) : colimit.ι ((OpenNhds.inclusion x).op ⋙ (smoothSheaf IM I M R).presheaf) U ≫ @@ -315,12 +315,13 @@ def smoothSheafCommRing.forgetStalk (x : TopCat.of M) : (colimit.ι ((OpenNhds.inclusion x).op ⋙ (smoothSheafCommRing IM I M R).presheaf) U) := by rw [Iso.comp_inv_eq, ← smoothSheafCommRing.ι_forgetStalk_hom, CommRingCat.forget_map] simp_rw [Functor.comp_obj, Functor.op_obj] + rfl /-- Given a smooth commutative ring `R` and a manifold `M`, and an open neighbourhood `U` of a point `x : M`, the evaluation-at-`x` map to `R` from smooth functions from `U` to `R`. -/ def smoothSheafCommRing.evalAt (x : TopCat.of M) (U : OpenNhds x) : (smoothSheafCommRing IM I M R).presheaf.obj (Opposite.op U.1) ⟶ CommRingCat.of R := - SmoothMap.evalRingHom ⟨x, U.2⟩ + CommRingCat.ofHom (SmoothMap.evalRingHom ⟨x, U.2⟩) /-- Canonical ring homomorphism from the stalk of `smoothSheafCommRing IM I M R` at `x` to `R`, given by evaluating sections at `x`, considered as a morphism in the category of commutative rings. @@ -334,7 +335,7 @@ def smoothSheafCommRing.evalHom (x : TopCat.of M) : /-- Canonical ring homomorphism from the stalk of `smoothSheafCommRing IM I M R` at `x` to `R`, given by evaluating sections at `x`. -/ def smoothSheafCommRing.eval (x : M) : (smoothSheafCommRing IM I M R).presheaf.stalk x →+* R := - smoothSheafCommRing.evalHom IM I M R x + (smoothSheafCommRing.evalHom IM I M R x).hom @[simp, reassoc, elementwise] lemma smoothSheafCommRing.ι_evalHom (x : TopCat.of M) (U) : colimit.ι ((OpenNhds.inclusion x).op ⋙ _) U ≫ smoothSheafCommRing.evalHom IM I M R x = @@ -351,7 +352,7 @@ def smoothSheafCommRing.eval (x : M) : (smoothSheafCommRing IM I M R).presheaf.s @[simp, reassoc, elementwise] lemma smoothSheafCommRing.forgetStalk_inv_comp_eval (x : TopCat.of M) : (smoothSheafCommRing.forgetStalk IM I M R x).inv ≫ - (DFunLike.coe (smoothSheafCommRing.evalHom IM I M R x)) = + (DFunLike.coe (smoothSheafCommRing.evalHom IM I M R x).hom) = smoothSheaf.evalHom _ _ _ _ := by apply Limits.colimit.hom_ext intro U diff --git a/Mathlib/Geometry/RingedSpace/Basic.lean b/Mathlib/Geometry/RingedSpace/Basic.lean index ec7eb24b857e9..69d6ef0396197 100644 --- a/Mathlib/Geometry/RingedSpace/Basic.lean +++ b/Mathlib/Geometry/RingedSpace/Basic.lean @@ -44,8 +44,9 @@ open SheafedSpace @[simp] lemma res_zero {X : RingedSpace.{u}} {U V : TopologicalSpace.Opens X} - (hUV : U ≤ V) : (0 : X.presheaf.obj (op V)) |_ U = 0 := - map_zero _ + (hUV : U ≤ V) : (0 : X.presheaf.obj (op V)) |_ U = + (0 : (CategoryTheory.forget CommRingCat).obj (X.presheaf.obj (op U))) := + RingHom.map_zero _ variable (X : RingedSpace) @@ -58,7 +59,8 @@ lemma exists_res_eq_zero_of_germ_eq_zero (U : Opens X) (f : X.presheaf.obj (op U (h : X.presheaf.germ U x.val x.property f = 0) : ∃ (V : Opens X) (i : V ⟶ U) (_ : x.1 ∈ V), X.presheaf.map i.op f = 0 := by have h1 : X.presheaf.germ U x.val x.property f = X.presheaf.germ U x.val x.property 0 := by simpa - obtain ⟨V, hv, i, _, hv4⟩ := TopCat.Presheaf.germ_eq X.presheaf x.1 x.2 x.2 f 0 h1 + obtain ⟨V, hv, i, _, (hv4 : (X.presheaf.map i.op) f = (X.presheaf.map _) 0)⟩ := + TopCat.Presheaf.germ_eq X.presheaf x.1 x.2 x.2 f 0 h1 use V, i, hv simpa using hv4 @@ -83,11 +85,36 @@ theorem isUnit_res_of_isUnit_germ (U : Opens X) (f : X.presheaf.obj (op U)) (x : show X.presheaf.germ _ x hxW (X.presheaf.map (U.infLERight V).op g) = X.presheaf.germ _ x hxV g from X.presheaf.germ_res_apply (Opens.infLERight U V) x hxW g] exact heq - obtain ⟨W', hxW', i₁, i₂, heq'⟩ := X.presheaf.germ_eq x hxW hxW _ _ heq + -- note: we have to force lean to resynthesize this as <...>.hom _ = <...>.hom _ + obtain ⟨W', hxW', i₁, i₂, (heq' : (X.presheaf.map i₁.op) _ = (X.presheaf.map i₂.op) 1)⟩ := + X.presheaf.germ_eq x hxW hxW _ _ heq use W', i₁ ≫ Opens.infLELeft U V, hxW' - rw [(X.presheaf.map i₂.op).map_one, (X.presheaf.map i₁.op).map_mul] at heq' - rw [← comp_apply, ← X.presheaf.map_comp, ← comp_apply, ← X.presheaf.map_comp, ← op_comp] at heq' - exact isUnit_of_mul_eq_one _ _ heq' + simp only [map_mul, map_one] at heq' + simpa using isUnit_of_mul_eq_one _ _ heq' + +/-- Specialize `TopCat.Presheaf.germ_res_apply` to sheaves of rings. + +This is unfortunately needed because the results on presheaves are stated using the +`ConcreteCategory.instFunLike` instance, which is not reducibly equal to the actual coercion of +morphisms in `CommRingCat` to functions. +-/ +lemma _root_.CommRingCat.germ_res_apply + {X : TopCat} (F : Presheaf CommRingCat X) + {U V : Opens X} (i : U ⟶ V) (x : X) (hx : x ∈ U) (s) : + F.germ U x hx (F.map i.op s) = F.germ V x (i.le hx) s := + F.germ_res_apply _ _ _ _ + +/-- Specialize `TopCat.Presheaf.germ_res_apply'` to sheaves of rings. + +This is unfortunately needed because the results on presheaves are stated using the +`ConcreteCategory.instFunLike` instance, which is not reducibly equal to the actual coercion of +morphisms in `CommRingCat` to functions. +-/ +lemma _root_.CommRingCat.germ_res_apply' + {X : TopCat} (F : Presheaf CommRingCat X) + {U V : Opens X} (i : op V ⟶ op U) (x : X) (hx : x ∈ U) (s) : + F.germ U x hx (F.map i s) = F.germ V x (i.unop.le hx) s := + F.germ_res_apply' _ _ _ _ /-- If a section `f` is a unit in each stalk, `f` must be a unit. -/ theorem isUnit_of_isUnit_germ (U : Opens X) (f : X.presheaf.obj (op U)) @@ -108,21 +135,28 @@ theorem isUnit_of_isUnit_germ (U : Opens X) (f : X.presheaf.obj (op U)) rw [germ_res_apply, germ_res_apply] apply (h z ((iVU x).le hzVx)).mul_right_inj.mp -- Porting note: now need explicitly typing the rewrites - rw [← X.presheaf.germ_res_apply (iVU x) z hzVx f] + -- note: this is bad, I think we should replace the `FunLike` on + -- concrete category with `CoeFun` + rw [← CommRingCat.germ_res_apply X.presheaf (iVU x) z hzVx f] -- Porting note: change was not necessary in Lean3 change X.presheaf.germ _ z hzVx _ * (X.presheaf.germ _ z hzVx _) = X.presheaf.germ _ z hzVx _ * X.presheaf.germ _ z hzVy (g y) rw [← RingHom.map_mul, congr_arg (X.presheaf.germ (V x) z hzVx) (hg x), - X.presheaf.germ_res_apply _ _ _ f, - ← X.presheaf.germ_res_apply (iVU y) z hzVy f, + CommRingCat.germ_res_apply X.presheaf _ _ _ f, + ← CommRingCat.germ_res_apply X.presheaf (iVU y) z hzVy f, ← RingHom.map_mul, congr_arg (X.presheaf.germ (V y) z hzVy) (hg y), RingHom.map_one, RingHom.map_one] -- We claim that these local inverses glue together to a global inverse of `f`. - obtain ⟨gl, gl_spec, -⟩ := X.sheaf.existsUnique_gluing' V U iVU hcover g ic + obtain ⟨gl, gl_spec, -⟩ : + -- We need to rephrase the result from `ConcreteCategory` to `CommRingCat`. + ∃ gl : X.presheaf.obj (op U), (∀ i, ((sheaf X).val.map (iVU i).op) gl = g i) ∧ _ := + X.sheaf.existsUnique_gluing' V U iVU hcover g ic apply isUnit_of_mul_eq_one f gl apply X.sheaf.eq_of_locally_eq' V U iVU hcover intro i + -- We need to rephrase the goal from `ConcreteCategory` to `CommRingCat`. + show ((sheaf X).val.map (iVU i).op).hom (f * gl) = ((sheaf X).val.map (iVU i).op) 1 rw [RingHom.map_one, RingHom.map_mul, gl_spec] exact hg i @@ -139,7 +173,7 @@ def basicOpen {U : Opens X} (f : X.presheaf.obj (op U)) : Opens X where refine ⟨?_, V.2, hxV⟩ intro y hy use i.le hy - convert RingHom.isUnit_map (X.presheaf.germ _ y hy) hf + convert RingHom.isUnit_map (X.presheaf.germ _ y hy).hom hf exact (X.presheaf.germ_res_apply i y hy f).symm theorem mem_basicOpen {U : Opens X} (f : X.presheaf.obj (op U)) (x : X) (hx : x ∈ U) : @@ -166,18 +200,18 @@ theorem isUnit_res_basicOpen {U : Opens X} (f : X.presheaf.obj (op U)) : apply isUnit_of_isUnit_germ rintro x ⟨hxU, hx⟩ convert hx - convert X.presheaf.germ_res_apply _ _ _ _ + exact X.presheaf.germ_res_apply _ _ _ _ @[simp] theorem basicOpen_res {U V : (Opens X)ᵒᵖ} (i : U ⟶ V) (f : X.presheaf.obj U) : @basicOpen X (unop V) (X.presheaf.map i f) = unop V ⊓ @basicOpen X (unop U) f := by ext x; constructor · rintro ⟨hxV, hx⟩ - rw [X.presheaf.germ_res_apply'] at hx + rw [CommRingCat.germ_res_apply' X.presheaf] at hx exact ⟨hxV, i.unop.le hxV, hx⟩ · rintro ⟨hxV, _, hx⟩ refine ⟨hxV, ?_⟩ - rw [X.presheaf.germ_res_apply'] + rw [CommRingCat.germ_res_apply' X.presheaf] exact hx -- This should fire before `basicOpen_res`. @@ -190,7 +224,8 @@ theorem basicOpen_res_eq {U V : (Opens X)ᵒᵖ} (i : U ⟶ V) [IsIso i] (f : X. apply le_antisymm · rw [X.basicOpen_res i f]; exact inf_le_right · have := X.basicOpen_res (inv i) (X.presheaf.map i f) - rw [← comp_apply, ← X.presheaf.map_comp, IsIso.hom_inv_id, X.presheaf.map_id, id_apply] at this + rw [← CommRingCat.comp_apply, ← X.presheaf.map_comp, IsIso.hom_inv_id, X.presheaf.map_id, + CommRingCat.id_apply] at this rw [this] exact inf_le_right diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean index ce60b4798409b..36647bd9a55ca 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean @@ -75,7 +75,7 @@ def 𝒪 : Sheaf CommRingCat X.toTopCat := structure Hom (X Y : LocallyRingedSpace.{u}) extends X.toPresheafedSpace.Hom Y.toPresheafedSpace : Type _ where /-- the underlying morphism induces a local ring homomorphism on stalks -/ - prop : ∀ x, IsLocalHom (toHom.stalkMap x) + prop : ∀ x, IsLocalHom (toHom.stalkMap x).hom /-- A morphism of locally ringed spaces as a morphism of sheafed spaces. -/ abbrev Hom.toShHom {X Y : LocallyRingedSpace.{u}} (f : X.Hom Y) : @@ -108,12 +108,12 @@ noncomputable def Hom.stalkMap {X Y : LocallyRingedSpace.{u}} (f : Hom X Y) (x : @[instance] theorem isLocalHomStalkMap {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) : - IsLocalHom (f.stalkMap x) := + IsLocalHom (f.stalkMap x).hom := f.2 x @[instance] theorem isLocalHomValStalkMap {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) : - IsLocalHom (f.toShHom.stalkMap x) := + IsLocalHom (f.toShHom.stalkMap x).hom := f.2 x @[deprecated (since := "2024-10-10")] @@ -122,7 +122,7 @@ alias isLocalRingHomValStalkMap := isLocalHomValStalkMap /-- The identity morphism on a locally ringed space. -/ @[simps! toShHom] def id (X : LocallyRingedSpace.{u}) : Hom X X := - ⟨𝟙 X.toSheafedSpace, fun x => by erw [PresheafedSpace.stalkMap.id]; infer_instance⟩ + ⟨𝟙 X.toSheafedSpace, fun x => by dsimp; erw [PresheafedSpace.stalkMap.id]; infer_instance⟩ instance (X : LocallyRingedSpace.{u}) : Inhabited (Hom X X) := ⟨id X⟩ @@ -130,6 +130,7 @@ instance (X : LocallyRingedSpace.{u}) : Inhabited (Hom X X) := /-- Composition of morphisms of locally ringed spaces. -/ def comp {X Y Z : LocallyRingedSpace.{u}} (f : Hom X Y) (g : Hom Y Z) : Hom X Z := ⟨f.toShHom ≫ g.toShHom, fun x => by + dsimp erw [PresheafedSpace.stalkMap.comp] infer_instance⟩ @@ -271,13 +272,13 @@ def empty : LocallyRingedSpace.{u} where instance : EmptyCollection LocallyRingedSpace.{u} := ⟨LocallyRingedSpace.empty⟩ /-- The canonical map from the empty locally ringed space. -/ -def emptyTo (X : LocallyRingedSpace) : ∅ ⟶ X := +def emptyTo (X : LocallyRingedSpace.{u}) : ∅ ⟶ X := ⟨⟨⟨fun x => PEmpty.elim x, by fun_prop⟩, - { app := fun U => by refine ⟨⟨⟨0, ?_⟩, ?_⟩, ?_, ?_⟩ <;> intros <;> rfl }⟩, + { app := fun U => CommRingCat.ofHom <| by refine ⟨⟨⟨0, ?_⟩, ?_⟩, ?_, ?_⟩ <;> intros <;> rfl }⟩, fun x => PEmpty.elim x⟩ noncomputable -instance {X : LocallyRingedSpace} : Unique (∅ ⟶ X) where +instance {X : LocallyRingedSpace.{u}} : Unique (∅ ⟶ X) where default := LocallyRingedSpace.emptyTo X uniq f := by ext ⟨⟩ x; aesop_cat @@ -312,7 +313,7 @@ lemma basicOpen_eq_bot_of_isNilpotent (X : LocallyRingedSpace.{u}) (U : Opens X. instance component_nontrivial (X : LocallyRingedSpace.{u}) (U : Opens X.carrier) [hU : Nonempty U] : Nontrivial (X.presheaf.obj <| op U) := - (X.presheaf.germ _ _ hU.some.2).domain_nontrivial + (X.presheaf.germ _ _ hU.some.2).hom.domain_nontrivial @[simp] lemma iso_hom_base_inv_base {X Y : LocallyRingedSpace.{u}} (e : X ≅ Y) : @@ -360,7 +361,7 @@ lemma stalkSpecializes_stalkMap (x x' : X) (h : x ⤳ x') : lemma stalkSpecializes_stalkMap_apply (x x' : X) (h : x ⤳ x') (y) : f.stalkMap x (Y.presheaf.stalkSpecializes (f.base.map_specializes h) y) = (X.presheaf.stalkSpecializes h (f.stalkMap x' y)) := - DFunLike.congr_fun (stalkSpecializes_stalkMap f x x' h) y + DFunLike.congr_fun (CommRingCat.hom_ext_iff.mp (stalkSpecializes_stalkMap f x x' h)) y @[reassoc] lemma stalkMap_congr (f g : X ⟶ Y) (hfg : f = g) (x x' : X) (hxx' : x = x') : @@ -395,7 +396,7 @@ lemma stalkMap_hom_inv (e : X ≅ Y) (y : Y) : lemma stalkMap_hom_inv_apply (e : X ≅ Y) (y : Y) (z) : e.inv.stalkMap y (e.hom.stalkMap (e.inv.base y) z) = Y.presheaf.stalkSpecializes (specializes_of_eq <| by simp) z := - DFunLike.congr_fun (stalkMap_hom_inv e y) z + DFunLike.congr_fun (CommRingCat.hom_ext_iff.mp (stalkMap_hom_inv e y)) z @[reassoc (attr := simp)] lemma stalkMap_inv_hom (e : X ≅ Y) (x : X) : @@ -408,7 +409,7 @@ lemma stalkMap_inv_hom (e : X ≅ Y) (x : X) : lemma stalkMap_inv_hom_apply (e : X ≅ Y) (x : X) (y) : e.hom.stalkMap x (e.inv.stalkMap (e.hom.base x) y) = X.presheaf.stalkSpecializes (specializes_of_eq <| by simp) y := - DFunLike.congr_fun (stalkMap_inv_hom e x) y + DFunLike.congr_fun (CommRingCat.hom_ext_iff.mp (stalkMap_inv_hom e x)) y @[reassoc (attr := simp)] lemma stalkMap_germ (U : Opens Y) (x : X) (hx : f.base x ∈ U) : @@ -431,12 +432,12 @@ theorem preimage_basicOpen {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) {U : Ope rw [SetLike.mem_coe, X.toRingedSpace.mem_basicOpen _ _ hxU] delta toRingedSpace rw [← stalkMap_germ_apply] - exact (f.stalkMap _).isUnit_map hx + exact (f.stalkMap _).hom.isUnit_map hx · rintro ⟨hxU, hx⟩ simp only [Opens.map_coe, Set.mem_preimage, SetLike.mem_coe, toRingedSpace] at hx ⊢ rw [RingedSpace.mem_basicOpen _ s (f.base x) hxU] rw [← stalkMap_germ_apply] at hx - exact (isUnit_map_iff (f.toShHom.stalkMap _) _).mp hx + exact (isUnit_map_iff (f.toShHom.stalkMap _).hom _).mp hx variable {U : TopCat} (X : LocallyRingedSpace.{u}) {f : U ⟶ X.toTopCat} (h : IsOpenEmbedding f) (V : Opens U) (x : U) (hx : x ∈ V) diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean index e8f22b8f1f79b..edc1b664e2ba8 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean @@ -103,7 +103,7 @@ noncomputable def coproductCofanIsColimit : IsColimit (coproductCofan F) where (forgetToSheafedSpace.mapCocone s) i : _)] haveI : IsLocalHom - (((forgetToSheafedSpace.mapCocone s).ι.app i).stalkMap y) := + (((forgetToSheafedSpace.mapCocone s).ι.app i).stalkMap y).hom := (s.ι.app i).2 y infer_instance⟩ fac _ _ := LocallyRingedSpace.Hom.ext' @@ -134,7 +134,7 @@ namespace HasCoequalizer @[instance] theorem coequalizer_π_app_isLocalHom (U : TopologicalSpace.Opens (coequalizer f.toShHom g.toShHom).carrier) : - IsLocalHom ((coequalizer.π f.toShHom g.toShHom : _).c.app (op U)) := by + IsLocalHom ((coequalizer.π f.toShHom g.toShHom : _).c.app (op U)).hom := by have := ι_comp_coequalizerComparison f.toShHom g.toShHom SheafedSpace.forgetToPresheafedSpace rw [← PreservesCoequalizer.iso_hom] at this erw [SheafedSpace.congr_app this.symm (op U)] @@ -143,6 +143,12 @@ theorem coequalizer_π_app_isLocalHom haveI : IsIso (PreservesCoequalizer.iso SheafedSpace.forgetToPresheafedSpace f.toShHom g.toShHom).hom.c := PresheafedSpace.c_isIso_of_iso _ + -- Had to add this instance too. + have := CommRingCat.equalizer_ι_is_local_ring_hom' (PresheafedSpace.componentwiseDiagram _ + ((Opens.map + (PreservesCoequalizer.iso SheafedSpace.forgetToPresheafedSpace (Hom.toShHom f) + (Hom.toShHom g)).hom.base).obj + (unop (op U)))) infer_instance @[deprecated (since := "2024-10-10")] @@ -196,9 +202,10 @@ theorem imageBasicOpen_image_preimage : rw [preimage_basicOpen f, preimage_basicOpen g] dsimp only [Functor.op, unop_op] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11224): change `rw` to `erw` - erw [← comp_apply, ← SheafedSpace.comp_c_app', ← comp_apply, ← SheafedSpace.comp_c_app', - SheafedSpace.congr_app (coequalizer.condition f.toShHom g.toShHom), comp_apply, - X.toRingedSpace.basicOpen_res] + erw [← CommRingCat.comp_apply, ← SheafedSpace.comp_c_app', ← CommRingCat.comp_apply, + ← SheafedSpace.comp_c_app', + SheafedSpace.congr_app (coequalizer.condition f.toShHom g.toShHom), + CommRingCat.comp_apply, X.toRingedSpace.basicOpen_res] apply inf_eq_right.mpr refine (RingedSpace.basicOpen_le _ _).trans ?_ rw [coequalizer.condition f.toShHom g.toShHom] @@ -216,12 +223,14 @@ theorem imageBasicOpen_image_open : @[instance] theorem coequalizer_π_stalk_isLocalHom (x : Y) : - IsLocalHom ((coequalizer.π f.toShHom g.toShHom : _).stalkMap x) := by + IsLocalHom ((coequalizer.π f.toShHom g.toShHom : _).stalkMap x).hom := by constructor rintro a ha - rcases TopCat.Presheaf.germ_exist _ _ a with ⟨U, hU, s, rfl⟩ - rw [PresheafedSpace.stalkMap_germ_apply + rcases TopCat.Presheaf.germ_exist (C := CommRingCat) _ _ a with ⟨U, hU, s, rfl⟩ + -- need `erw` to see through `ConcreteCategory.instFunLike` + rw [← CommRingCat.forget_map_apply, PresheafedSpace.stalkMap_germ_apply (coequalizer.π (C := SheafedSpace _) f.toShHom g.toShHom) U _ hU] at ha + rw [CommRingCat.forget_map_apply] let V := imageBasicOpen f g U s have hV : (coequalizer.π f.toShHom g.toShHom).base ⁻¹' ((coequalizer.π f.toShHom g.toShHom).base '' V.1) = V.1 := @@ -234,14 +243,14 @@ theorem coequalizer_π_stalk_isLocalHom (x : Y) : have VleU : (⟨(coequalizer.π f.toShHom g.toShHom).base '' V.1, V_open⟩ : _) ≤ U := Set.image_subset_iff.mpr (Y.toRingedSpace.basicOpen_le _) have hxV : x ∈ V := ⟨hU, ha⟩ - rw [← - (coequalizer f.toShHom g.toShHom).presheaf.germ_res_apply (homOfLE VleU) _ + rw [← CommRingCat.germ_res_apply (coequalizer f.toShHom g.toShHom).presheaf (homOfLE VleU) _ (@Set.mem_image_of_mem _ _ (coequalizer.π f.toShHom g.toShHom).base x V.1 hxV) s] apply RingHom.isUnit_map - rw [← isUnit_map_iff ((coequalizer.π f.toShHom g.toShHom : _).c.app _), ← comp_apply, - NatTrans.naturality, comp_apply, ← isUnit_map_iff (Y.presheaf.map (eqToHom hV').op)] + rw [← isUnit_map_iff ((coequalizer.π f.toShHom g.toShHom : _).c.app _).hom, + ← CommRingCat.comp_apply, NatTrans.naturality, CommRingCat.comp_apply, + ← isUnit_map_iff (Y.presheaf.map (eqToHom hV').op).hom] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11224): change `rw` to `erw` - erw [← comp_apply, ← comp_apply, ← Y.presheaf.map_comp] + erw [← CommRingCat.comp_apply, ← CommRingCat.comp_apply, ← Y.presheaf.map_comp] convert @RingedSpace.isUnit_res_basicOpen Y.toRingedSpace (unop _) (((coequalizer.π f.toShHom g.toShHom).c.app (op U)) s) @@ -257,7 +266,7 @@ noncomputable def coequalizer : LocallyRingedSpace where isLocalRing x := by obtain ⟨y, rfl⟩ := (TopCat.epi_iff_surjective (coequalizer.π f.toShHom g.toShHom).base).mp inferInstance x - exact ((coequalizer.π f.toShHom g.toShHom : _).stalkMap y).domain_isLocalRing + exact ((coequalizer.π f.toShHom g.toShHom : _).stalkMap y).hom.domain_isLocalRing /-- The explicit coequalizer cofork of locally ringed spaces. -/ noncomputable def coequalizerCofork : Cofork f g := @@ -267,8 +276,8 @@ noncomputable def coequalizerCofork : Cofork f g := (LocallyRingedSpace.Hom.ext' (coequalizer.condition f.toShHom g.toShHom)) theorem isLocalHom_stalkMap_congr {X Y : RingedSpace} (f g : X ⟶ Y) (H : f = g) (x) - (h : IsLocalHom (f.stalkMap x)) : - IsLocalHom (g.stalkMap x) := by + (h : IsLocalHom (f.stalkMap x).hom) : + IsLocalHom (g.stalkMap x).hom := by rw [PresheafedSpace.stalkMap.congr_hom _ _ H.symm x]; infer_instance /-- The cofork constructed in `coequalizer_cofork` is indeed a colimit cocone. -/ @@ -285,12 +294,9 @@ noncomputable def coequalizerCoforkIsColimit : IsColimit (coequalizerCofork f g) -- but this is no longer possible set h := _ change IsLocalHom h - suffices _ : IsLocalHom (((coequalizerCofork f g).π.1.stalkMap _).comp h) by - apply isLocalHom_of_comp _ ((coequalizerCofork f g).π.1.stalkMap _) - -- note to reviewers: this `change` is now more brittle because it now has to fully resolve - -- the type to be able to search for `MonoidHomClass`, even though of course all homs in - -- `CommRingCat` are clearly such - change IsLocalHom (h ≫ (coequalizerCofork f g).π.toShHom.stalkMap y) + suffices _ : IsLocalHom (((coequalizerCofork f g).π.1.stalkMap _).hom.comp h) by + apply isLocalHom_of_comp _ ((coequalizerCofork f g).π.1.stalkMap _).hom + rw [← CommRingCat.hom_ofHom h, ← CommRingCat.hom_comp] erw [← PresheafedSpace.stalkMap.comp] apply isLocalHom_stalkMap_congr _ _ (coequalizer.π_desc s.π.toShHom e).symm y infer_instance diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean index a5e9107f3709e..e96b86ffdeceb 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean @@ -52,7 +52,9 @@ If we interpret sections over `U` as functions of `X` defined on `U`, then this corresponds to evaluation at `x`. -/ def evaluation (x : U) : X.presheaf.obj (op U) ⟶ X.residueField x := - X.presheaf.germ U x.1 x.2 ≫ IsLocalRing.residue _ + -- TODO: make a new definition wrapping + -- `CommRingCat.ofHom (IsLocalRing.residue (X.presheaf.stalk _))`? + X.presheaf.germ U x.1 x.2 ≫ CommRingCat.ofHom (IsLocalRing.residue (X.presheaf.stalk _)) /-- The global evaluation map from `Γ(X, ⊤)` to the residue field at `x`. -/ def Γevaluation (x : X) : X.presheaf.obj (op ⊤) ⟶ X.residueField x := @@ -88,23 +90,28 @@ variable {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) /-- If `X ⟶ Y` is a morphism of locally ringed spaces and `x` a point of `X`, we obtain a morphism of residue fields in the other direction. -/ def residueFieldMap (x : X) : Y.residueField (f.base x) ⟶ X.residueField x := - IsLocalRing.ResidueField.map (f.stalkMap x) + CommRingCat.ofHom (IsLocalRing.ResidueField.map (f.stalkMap x).hom) lemma residue_comp_residueFieldMap_eq_stalkMap_comp_residue (x : X) : - IsLocalRing.residue _ ≫ residueFieldMap f x = f.stalkMap x ≫ IsLocalRing.residue _ := by + CommRingCat.ofHom (IsLocalRing.residue (Y.presheaf.stalk (f.base x))) ≫ + residueFieldMap f x = f.stalkMap x ≫ + CommRingCat.ofHom (IsLocalRing.residue (X.presheaf.stalk x)) := by simp [residueFieldMap] rfl @[simp] lemma residueFieldMap_id (x : X) : residueFieldMap (𝟙 X) x = 𝟙 (X.residueField x) := by + ext : 1 simp only [id_toShHom', SheafedSpace.id_base, TopCat.coe_id, id_eq, residueFieldMap, stalkMap_id] apply IsLocalRing.ResidueField.map_id @[simp] lemma residueFieldMap_comp {Z : LocallyRingedSpace.{u}} (g : Y ⟶ Z) (x : X) : residueFieldMap (f ≫ g) x = residueFieldMap g (f.base x) ≫ residueFieldMap f x := by - simp only [comp_toShHom, SheafedSpace.comp_base, Function.comp_apply, residueFieldMap] + ext : 1 + simp only [comp_toShHom, SheafedSpace.comp_base, Function.comp_apply, residueFieldMap, + CommRingCat.hom_comp, TopCat.comp_app] simp_rw [stalkMap_comp] apply IsLocalRing.ResidueField.map_comp @@ -116,15 +123,17 @@ lemma evaluation_naturality {V : Opens Y} (x : (Opens.map f.base).obj V) : LocallyRingedSpace.residueFieldMap] rw [Category.assoc] ext a - simp only [comp_apply] - erw [IsLocalRing.ResidueField.map_residue, PresheafedSpace.stalkMap_germ_apply] + simp only [CommRingCat.comp_apply] + erw [IsLocalRing.ResidueField.map_residue] + rw [LocallyRingedSpace.stalkMap_germ_apply] rfl lemma evaluation_naturality_apply {V : Opens Y} (x : (Opens.map f.base).obj V) (a : Y.presheaf.obj (op V)) : residueFieldMap f x.val (Y.evaluation ⟨f.base x, x.property⟩ a) = X.evaluation x (f.c.app (op V) a) := by - simpa using congrFun (congrArg DFunLike.coe <| evaluation_naturality f x) a + simpa using congrFun (congrArg (DFunLike.coe ∘ CommRingCat.Hom.hom) <| + evaluation_naturality f x) a @[reassoc] lemma Γevaluation_naturality (x : X) : diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index d2c8e5bec1786..697a19f7df281 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -964,6 +964,7 @@ def pullbackConeOfLeft : PullbackCone f g := by rw [PresheafedSpace.stalkMap.comp, PresheafedSpace.stalkMap.comp] at this rw [← IsIso.eq_inv_comp] at this rw [this] + dsimp infer_instance · exact LocallyRingedSpace.Hom.ext' (PresheafedSpace.IsOpenImmersion.pullback_cone_of_left_condition _ _) @@ -1206,7 +1207,7 @@ theorem invApp_app (U : Opens X) : (eqToHom (by simp [Opens.map, Set.preimage_image_eq _ H.base_open.injective])) := PresheafedSpace.IsOpenImmersion.invApp_app f.1 U -attribute [elementwise] invApp_app +attribute [elementwise nosimp] invApp_app @[reassoc (attr := simp)] theorem app_invApp (U : Opens Y) : diff --git a/Mathlib/RingTheory/RingHomProperties.lean b/Mathlib/RingTheory/RingHomProperties.lean index 20912a43d1f3b..3f6caf02c496d 100644 --- a/Mathlib/RingTheory/RingHomProperties.lean +++ b/Mathlib/RingTheory/RingHomProperties.lean @@ -45,17 +45,17 @@ def RespectsIso : Prop := variable {P} theorem RespectsIso.cancel_left_isIso (hP : RespectsIso @P) {R S T : CommRingCat} (f : R ⟶ S) - (g : S ⟶ T) [IsIso f] : P (f ≫ g) ↔ P g := + (g : S ⟶ T) [IsIso f] : P (g.hom.comp f.hom) ↔ P g.hom := ⟨fun H => by - convert hP.2 (f ≫ g) (asIso f).symm.commRingCatIsoToRingEquiv H - exact (IsIso.inv_hom_id_assoc _ _).symm, hP.2 g (asIso f).commRingCatIsoToRingEquiv⟩ + convert hP.2 (f ≫ g).hom (asIso f).symm.commRingCatIsoToRingEquiv H + exact (IsIso.inv_hom_id_assoc _ _).symm, hP.2 g.hom (asIso f).commRingCatIsoToRingEquiv⟩ theorem RespectsIso.cancel_right_isIso (hP : RespectsIso @P) {R S T : CommRingCat} (f : R ⟶ S) - (g : S ⟶ T) [IsIso g] : P (f ≫ g) ↔ P f := + (g : S ⟶ T) [IsIso g] : P (g.hom.comp f.hom) ↔ P f.hom := ⟨fun H => by - convert hP.1 (f ≫ g) (asIso g).symm.commRingCatIsoToRingEquiv H + convert hP.1 (f ≫ g).hom (asIso g).symm.commRingCatIsoToRingEquiv H change f = f ≫ g ≫ inv g - simp, hP.1 f (asIso g).commRingCatIsoToRingEquiv⟩ + simp, hP.1 f.hom (asIso g).commRingCatIsoToRingEquiv⟩ theorem RespectsIso.is_localization_away_iff (hP : RingHom.RespectsIso @P) {R S : Type u} (R' S' : Type u) [CommRing R] [CommRing S] [CommRing R'] [CommRing S'] [Algebra R R'] @@ -156,13 +156,13 @@ theorem IsStableUnderBaseChange.mk (h₁ : RespectsIso @P) attribute [local instance] Algebra.TensorProduct.rightAlgebra theorem IsStableUnderBaseChange.pushout_inl (hP : RingHom.IsStableUnderBaseChange @P) - (hP' : RingHom.RespectsIso @P) {R S T : CommRingCat} (f : R ⟶ S) (g : R ⟶ T) (H : P g) : - P (pushout.inl _ _ : S ⟶ pushout f g) := by - letI := f.toAlgebra - letI := g.toAlgebra + (hP' : RingHom.RespectsIso @P) {R S T : CommRingCat} (f : R ⟶ S) (g : R ⟶ T) (H : P g.hom) : + P (pushout.inl _ _ : S ⟶ pushout f g).hom := by + letI := f.hom.toAlgebra + letI := g.hom.toAlgebra rw [← show _ = pushout.inl f g from colimit.isoColimitCocone_ι_inv ⟨_, CommRingCat.pushoutCoconeIsColimit R S T⟩ WalkingSpan.left, - hP'.cancel_right_isIso] + CommRingCat.hom_comp, hP'.cancel_right_isIso] dsimp only [CommRingCat.pushoutCocone_inl, PushoutCocone.ι_app_left] apply hP R T S (TensorProduct R S T) exact H @@ -173,7 +173,7 @@ section ToMorphismProperty /-- The categorical `MorphismProperty` associated to a property of ring homs expressed non-categorical terms. -/ -def toMorphismProperty : MorphismProperty CommRingCat := fun _ _ f ↦ P f +def toMorphismProperty : MorphismProperty CommRingCat := fun _ _ f ↦ P f.hom variable {P} @@ -181,9 +181,9 @@ lemma toMorphismProperty_respectsIso_iff : RespectsIso P ↔ (toMorphismProperty P).RespectsIso := by refine ⟨fun h ↦ MorphismProperty.RespectsIso.mk _ ?_ ?_, fun h ↦ ⟨?_, ?_⟩⟩ · intro X Y Z e f hf - exact h.right f e.commRingCatIsoToRingEquiv hf + exact h.right f.hom e.commRingCatIsoToRingEquiv hf · intro X Y Z e f hf - exact h.left f e.commRingCatIsoToRingEquiv hf + exact h.left f.hom e.commRingCatIsoToRingEquiv hf · intro X Y Z _ _ _ f e hf exact MorphismProperty.RespectsIso.postcomp (toMorphismProperty P) e.toCommRingCatIso.hom (CommRingCat.ofHom f) hf diff --git a/Mathlib/Topology/Category/TopCommRingCat.lean b/Mathlib/Topology/Category/TopCommRingCat.lean index d8e2fce858d9a..27530487131f3 100644 --- a/Mathlib/Topology/Category/TopCommRingCat.lean +++ b/Mathlib/Topology/Category/TopCommRingCat.lean @@ -77,7 +77,8 @@ instance forgetTopologicalRing (R : TopCommRingCat) : R.isTopologicalRing instance hasForgetToCommRingCat : HasForget₂ TopCommRingCat CommRingCat := - HasForget₂.mk' (fun R => CommRingCat.of R) (fun _ => rfl) (fun f => f.val) HEq.rfl + HasForget₂.mk' (fun R => CommRingCat.of R) (fun _ => rfl) + (fun f => CommRingCat.ofHom f.val) HEq.rfl instance forgetToCommRingCatTopologicalSpace (R : TopCommRingCat) : TopologicalSpace ((forget₂ TopCommRingCat CommRingCat).obj R) := diff --git a/Mathlib/Topology/Sheaves/CommRingCat.lean b/Mathlib/Topology/Sheaves/CommRingCat.lean index d54fbeeae9415..0b46eb90f16d8 100644 --- a/Mathlib/Topology/Sheaves/CommRingCat.lean +++ b/Mathlib/Topology/Sheaves/CommRingCat.lean @@ -47,6 +47,38 @@ example (X : TopCat.{u₁}) (F : Presheaf CommRingCat.{u₁} X) F.IsSheaf := (isSheaf_iff_isSheaf_comp (forget CommRingCat) F).mpr h +/-- +Specialize `restrictOpen` to `CommRingCat` because inferring `C := CommRingCat` isn't reliable. +Instead of unfolding the definition, rewrite with `restrictOpenCommRingCat_apply` to ensure the +correct coercion to functions is taken. + +(The correct fix in the longer term is to redesign concrete categories so we don't use `forget` +everywhere, but the correct `FunLike` instance for the morphisms of those categories.) +-/ +abbrev restrictOpenCommRingCat {X : TopCat} + {F : Presheaf CommRingCat X} {V : Opens ↑X} (f : CommRingCat.carrier (F.obj (op V))) + (U : Opens ↑X) (e : U ≤ V := by restrict_tac) : + CommRingCat.carrier (F.obj (op U)) := + TopCat.Presheaf.restrictOpen (C := CommRingCat) f U e + +/-- Notation for `TopCat.Presheaf.restrictOpenCommRingCat`. -/ +scoped[AlgebraicGeometry] infixl:80 " |_ᵣ " => TopCat.Presheaf.restrictOpenCommRingCat + +open AlgebraicGeometry in +lemma restrictOpenCommRingCat_apply {X : TopCat} + {F : Presheaf CommRingCat X} {V : Opens ↑X} (f : CommRingCat.carrier (F.obj (op V))) + (U : Opens ↑X) (e : U ≤ V := by restrict_tac) : + f |_ᵣ U = F.map (homOfLE e).op f := + rfl + +open AlgebraicGeometry in +lemma _root_.CommRingCat.presheaf_restrict_restrict (X : TopCat) + {F : TopCat.Presheaf CommRingCat X} + {U V W : Opens ↑X} (e₁ : U ≤ V := by restrict_tac) (e₂ : V ≤ W := by restrict_tac) + (f : CommRingCat.carrier (F.obj (op W))) : + f |_ᵣ V |_ᵣ U = f |_ᵣ U := + TopCat.Presheaf.restrict_restrict (C := CommRingCat) e₁ e₂ f + section SubmonoidPresheaf open scoped nonZeroDivisors @@ -55,18 +87,18 @@ variable {X : TopCat.{w}} {C : Type u} [Category.{v} C] [ConcreteCategory C] attribute [local instance 1000] ConcreteCategory.hasCoeToSort ConcreteCategory.instFunLike +-- note: this was specialized to `CommRingCat` in #19757 /-- A subpresheaf with a submonoid structure on each of the components. -/ -structure SubmonoidPresheaf [∀ X : C, MulOneClass X] [∀ X Y : C, MonoidHomClass (X ⟶ Y) X Y] - (F : X.Presheaf C) where +structure SubmonoidPresheaf (F : X.Presheaf CommRingCat) where obj : ∀ U, Submonoid (F.obj U) - map : ∀ {U V : (Opens X)ᵒᵖ} (i : U ⟶ V), obj U ≤ (obj V).comap (F.map i) + map : ∀ {U V : (Opens X)ᵒᵖ} (i : U ⟶ V), obj U ≤ (obj V).comap (F.map i).hom variable {F : X.Presheaf CommRingCat.{w}} (G : F.SubmonoidPresheaf) /-- The localization of a presheaf of `CommRing`s with respect to a `SubmonoidPresheaf`. -/ protected noncomputable def SubmonoidPresheaf.localizationPresheaf : X.Presheaf CommRingCat where obj U := CommRingCat.of <| Localization (G.obj U) - map {_ _} i := CommRingCat.ofHom <| IsLocalization.map _ (F.map i) (G.map i) + map {_ _} i := CommRingCat.ofHom <| IsLocalization.map _ (F.map i).hom (G.map i) map_id U := by simp_rw [F.map_id] ext x @@ -74,11 +106,13 @@ protected noncomputable def SubmonoidPresheaf.localizationPresheaf : X.Presheaf exact IsLocalization.map_id (M := G.obj U) (S := Localization (G.obj U)) x map_comp {U V W} i j := by delta CommRingCat.ofHom CommRingCat.of Bundled.of - simp_rw [F.map_comp, CommRingCat.comp_eq_ring_hom_comp] + simp_rw [F.map_comp] + ext : 1 + dsimp rw [IsLocalization.map_comp_map] -- Porting note: this instance can't be synthesized -instance (U) : Algebra ((forget CommRingCat).obj (F.obj U)) (G.localizationPresheaf.obj U) := +instance (U) : Algebra (F.obj U) (G.localizationPresheaf.obj U) := show Algebra _ (Localization (G.obj U)) from inferInstance -- Porting note: this instance can't be synthesized @@ -89,7 +123,7 @@ instance (U) : IsLocalization (G.obj U) (G.localizationPresheaf.obj U) := @[simps app] def SubmonoidPresheaf.toLocalizationPresheaf : F ⟶ G.localizationPresheaf where app U := CommRingCat.ofHom <| algebraMap (F.obj U) (Localization <| G.obj U) - naturality {_ _} i := (IsLocalization.map_comp (G.map i)).symm + naturality {_ _} i := CommRingCat.hom_ext <| (IsLocalization.map_comp (G.map i)).symm instance epi_toLocalizationPresheaf : Epi G.toLocalizationPresheaf := @NatTrans.epi_of_epi_app _ _ _ _ _ _ G.toLocalizationPresheaf fun U => Localization.epi' (G.obj U) @@ -101,7 +135,7 @@ sections whose restriction onto each stalk falls in the given submonoid. -/ @[simps] noncomputable def submonoidPresheafOfStalk (S : ∀ x : X, Submonoid (F.stalk x)) : F.SubmonoidPresheaf where - obj U := ⨅ x : U.unop, Submonoid.comap (F.germ U.unop x.1 x.2) (S x) + obj U := ⨅ x : U.unop, Submonoid.comap (F.germ U.unop x.1 x.2).hom (S x) map {U V} i := by intro s hs simp only [Submonoid.mem_comap, Submonoid.mem_iInf] at hs ⊢ @@ -142,7 +176,8 @@ instance (F : X.Sheaf CommRingCat.{w}) : Mono F.presheaf.toTotalQuotientPresheaf intro s hs t e apply section_ext F (unop U) intro x hx - rw [map_zero] + show (F.presheaf.germ (unop U) x hx) t = (F.presheaf.germ (unop U) x hx) 0 + rw [RingHom.map_zero] apply Submonoid.mem_iInf.mp hs ⟨x, hx⟩ rw [← map_mul, e, map_zero] @@ -167,26 +202,30 @@ def continuousFunctions (X : TopCat.{v}ᵒᵖ) (R : TopCommRingCat.{v}) : CommRi namespace continuousFunctions +instance (X : TopCat.{v}ᵒᵖ) (R : TopCommRingCat.{v}) : + CommRing (unop X ⟶ (forget₂ TopCommRingCat TopCat).obj R) := + inferInstanceAs (CommRing (ContinuousMap _ _)) + /-- Pulling back functions into a topological ring along a continuous map is a ring homomorphism. -/ def pullback {X Y : TopCatᵒᵖ} (f : X ⟶ Y) (R : TopCommRingCat) : - continuousFunctions X R ⟶ continuousFunctions Y R where - toFun g := f.unop ≫ g - map_one' := rfl - map_zero' := rfl - map_add' := by aesop_cat - map_mul' := by aesop_cat + continuousFunctions X R ⟶ continuousFunctions Y R := CommRingCat.ofHom + { toFun g := f.unop ≫ g + map_one' := rfl + map_zero' := rfl + map_add' := by aesop_cat + map_mul' := by aesop_cat } /-- A homomorphism of topological rings can be postcomposed with functions from a source space `X`; this is a ring homomorphism (with respect to the pointwise ring operations on functions). -/ def map (X : TopCat.{u}ᵒᵖ) {R S : TopCommRingCat.{u}} (φ : R ⟶ S) : - continuousFunctions X R ⟶ continuousFunctions X S where - toFun g := g ≫ (forget₂ TopCommRingCat TopCat).map φ - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): `ext` tactic does not work, since Lean can't see through `R ⟶ S` is just - -- continuous ring homomorphism - map_one' := ContinuousMap.ext fun _ => φ.1.map_one - map_zero' := ContinuousMap.ext fun _ => φ.1.map_zero - map_add' := fun _ _ => ContinuousMap.ext fun _ => φ.1.map_add _ _ - map_mul' := fun _ _ => ContinuousMap.ext fun _ => φ.1.map_mul _ _ + continuousFunctions X R ⟶ continuousFunctions X S := CommRingCat.ofHom + { toFun g := g ≫ (forget₂ TopCommRingCat TopCat).map φ + -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): `ext` tactic does not work, since Lean can't see through `R ⟶ S` is just + -- continuous ring homomorphism + map_one' := ContinuousMap.ext fun _ => φ.1.map_one + map_zero' := ContinuousMap.ext fun _ => φ.1.map_zero + map_add' := fun _ _ => ContinuousMap.ext fun _ => φ.1.map_add _ _ + map_mul' := fun _ _ => ContinuousMap.ext fun _ => φ.1.map_mul _ _ } end continuousFunctions @@ -233,11 +272,11 @@ variable {X Y Z : TopCat.{v}} instance algebra_section_stalk (F : X.Presheaf CommRingCat) {U : Opens X} (x : U) : Algebra (F.obj <| op U) (F.stalk x) := - (F.germ U x.1 x.2).toAlgebra + (F.germ U x.1 x.2).hom.toAlgebra @[simp] theorem stalk_open_algebraMap {X : TopCat} (F : X.Presheaf CommRingCat) {U : Opens X} (x : U) : - algebraMap (F.obj <| op U) (F.stalk x) = F.germ U x.1 x.2 := + algebraMap (F.obj <| op U) (F.stalk x) = (F.germ U x.1 x.2).hom := rfl end TopCat.Presheaf @@ -261,9 +300,9 @@ def objSupIsoProdEqLocus {X : TopCat} (F : X.Sheaf CommRingCat) (U V : Opens X) F.1.obj (op <| U ⊔ V) ≅ CommRingCat.of <| -- Porting note: Lean 3 is able to figure out the ring homomorphism automatically RingHom.eqLocus - (RingHom.comp (F.val.map (homOfLE inf_le_left : U ⊓ V ⟶ U).op) + (RingHom.comp (F.val.map (homOfLE inf_le_left : U ⊓ V ⟶ U).op).hom (RingHom.fst (F.val.obj <| op U) (F.val.obj <| op V))) - (RingHom.comp (F.val.map (homOfLE inf_le_right : U ⊓ V ⟶ V).op) + (RingHom.comp (F.val.map (homOfLE inf_le_right : U ⊓ V ⟶ V).op).hom (RingHom.snd (F.val.obj <| op U) (F.val.obj <| op V))) := (F.isLimitPullbackCone U V).conePointUniqueUpToIso (CommRingCat.pullbackConeIsLimit _ _) @@ -308,20 +347,24 @@ theorem objSupIsoProdEqLocus_inv_eq_iff {X : TopCat.{u}} (F : X.Sheaf CommRingCa · rintro rfl rw [← TopCat.Sheaf.objSupIsoProdEqLocus_inv_fst, ← TopCat.Sheaf.objSupIsoProdEqLocus_inv_snd] -- `simp` doesn't see through the type equality of objects in `CommRingCat`, so use `rw` https://github.com/leanprover-community/mathlib4/pull/8386 - repeat rw [← comp_apply] + repeat rw [← CommRingCat.comp_apply] simp only [← Functor.map_comp, ← op_comp, Category.assoc, homOfLE_comp, and_self] · rintro ⟨e₁, e₂⟩ refine F.eq_of_locally_eq₂ (homOfLE (inf_le_right : U ⊓ W ≤ W)) (homOfLE (inf_le_right : V ⊓ W ≤ W)) ?_ _ _ ?_ ?_ · rw [← inf_sup_right] exact le_inf e le_rfl - · rw [← e₁, ← TopCat.Sheaf.objSupIsoProdEqLocus_inv_fst] + · change (F.val.map _) + ((F.val.map (homOfLE e).op).hom ((F.objSupIsoProdEqLocus U V).inv.hom x)) = (F.val.map _) y + rw [← e₁, ← TopCat.Sheaf.objSupIsoProdEqLocus_inv_fst] -- `simp` doesn't see through the type equality of objects in `CommRingCat`, so use `rw` https://github.com/leanprover-community/mathlib4/pull/8386 - repeat rw [← comp_apply] + repeat rw [← CommRingCat.comp_apply] simp only [← Functor.map_comp, ← op_comp, Category.assoc, homOfLE_comp] - · rw [← e₂, ← TopCat.Sheaf.objSupIsoProdEqLocus_inv_snd] + · show (F.val.map _) + ((F.val.map (homOfLE e).op).hom ((F.objSupIsoProdEqLocus U V).inv.hom x)) = (F.val.map _) y + rw [← e₂, ← TopCat.Sheaf.objSupIsoProdEqLocus_inv_snd] -- `simp` doesn't see through the type equality of objects in `CommRingCat`, so use `rw` https://github.com/leanprover-community/mathlib4/pull/8386 - repeat rw [← comp_apply] + repeat rw [← CommRingCat.comp_apply] simp only [← Functor.map_comp, ← op_comp, Category.assoc, homOfLE_comp] end TopCat.Sheaf diff --git a/MathlibTest/CategoryTheory/ConcreteCategory/Ring.lean b/MathlibTest/CategoryTheory/ConcreteCategory/Ring.lean new file mode 100644 index 0000000000000..a8f0352319c56 --- /dev/null +++ b/MathlibTest/CategoryTheory/ConcreteCategory/Ring.lean @@ -0,0 +1,53 @@ +import Mathlib.Algebra.Category.Ring.Basic + +universe v u + +open CategoryTheory SemiRingCat + +set_option maxHeartbeats 10000 +set_option synthInstance.maxHeartbeats 2000 + +/- We test if all the coercions and `map_add` lemmas trigger correctly. -/ + +example (X : Type u) [Semiring X] : ⇑(𝟙 (of X)) = id := by simp + +example {X Y : Type u} [Semiring X] [Semiring Y] (f : X →+* Y) : + ⇑(ofHom f) = ⇑f := by simp + +example {X Y : Type u} [Semiring X] [Semiring Y] (f : X →+* Y) + (x : X) : (ofHom f) x = f x := by simp + +example {X Y Z : SemiRingCat} (f : X ⟶ Y) (g : Y ⟶ Z) : ⇑(f ≫ g) = ⇑g ∘ ⇑f := by simp + +example {X Y Z : Type u} [Semiring X] [Semiring Y] [Semiring Z] + (f : X →+* Y) (g : Y →+* Z) : + ⇑(ofHom f ≫ ofHom g) = g ∘ f := by simp + +example {X Y : Type u} [Semiring X] [Semiring Y] {Z : SemiRingCat} + (f : X →+* Y) (g : of Y ⟶ Z) : + ⇑(ofHom f ≫ g) = g ∘ f := by simp + +example {X Y : SemiRingCat} {Z : Type u} [Semiring Z] (f : X ⟶ Y) (g : Y ⟶ of Z) : + ⇑(f ≫ g) = g ∘ f := by simp + +example {Y Z : SemiRingCat} {X : Type u} [Semiring X] (f : of X ⟶ Y) (g : Y ⟶ Z) : + ⇑(f ≫ g) = g ∘ f := by simp + +example {X Y Z : SemiRingCat} (f : X ⟶ Y) (g : Y ⟶ Z) (x : X) : (f ≫ g) x = g (f x) := by simp + +example {X Y : SemiRingCat} (e : X ≅ Y) (x : X) : e.inv (e.hom x) = x := by simp + +example {X Y : SemiRingCat} (e : X ≅ Y) (y : Y) : e.hom (e.inv y) = y := by simp + +example (X : SemiRingCat) : ⇑(𝟙 X) = id := by simp + +example {X : Type*} [Semiring X] : ⇑(RingHom.id X) = id := by simp + +example {M N : SemiRingCat} (f : M ⟶ N) (x y : M) : f (x + y) = f x + f y := by + simp + +example {M N : SemiRingCat} (f : M ⟶ N) (x y : M) : f (x * y) = f x * f y := by + simp + +example {M N : SemiRingCat} (f : M ⟶ N) : f 0 = 0 := by + simp From 98dbd12b96d0c47ff7292af3340fabc0d27dd125 Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Fri, 13 Dec 2024 15:22:26 +0000 Subject: [PATCH 701/829] feat(Algebra/Category): categories of modules have the same AB axioms as abelian groups (#19939) --- Mathlib.lean | 1 + Mathlib/Algebra/Category/ModuleCat/AB.lean | 32 ++++++++++++++++++++++ 2 files changed, 33 insertions(+) create mode 100644 Mathlib/Algebra/Category/ModuleCat/AB.lean diff --git a/Mathlib.lean b/Mathlib.lean index ed5502de95197..dfa70e2924e2d 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -92,6 +92,7 @@ import Mathlib.Algebra.Category.Grp.ZModuleEquivalence import Mathlib.Algebra.Category.Grp.Zero import Mathlib.Algebra.Category.GrpWithZero import Mathlib.Algebra.Category.HopfAlgebraCat.Basic +import Mathlib.Algebra.Category.ModuleCat.AB import Mathlib.Algebra.Category.ModuleCat.Abelian import Mathlib.Algebra.Category.ModuleCat.Adjunctions import Mathlib.Algebra.Category.ModuleCat.Algebra diff --git a/Mathlib/Algebra/Category/ModuleCat/AB.lean b/Mathlib/Algebra/Category/ModuleCat/AB.lean new file mode 100644 index 0000000000000..b6ab77266ed0d --- /dev/null +++ b/Mathlib/Algebra/Category/ModuleCat/AB.lean @@ -0,0 +1,32 @@ +/- +Copyright (c) 2024 Dagur Asgeirsson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Dagur Asgeirsson +-/ +import Mathlib.Algebra.Category.Grp.AB +import Mathlib.Algebra.Category.ModuleCat.Colimits +/-! + +# AB axioms in module categories + +This file proves that the category of modules over a ring satisfies Grothendieck's axioms AB5, AB4, +and AB4*. +-/ + +universe u + +open CategoryTheory Limits + +variable (R : Type u) [Ring R] + +instance : AB5 (ModuleCat.{u} R) where + ofShape J _ _ := + HasExactColimitsOfShape.domain_of_functor J (forget₂ (ModuleCat R) AddCommGrp) + +attribute [local instance] Abelian.hasFiniteBiproducts + +instance : AB4 (ModuleCat.{u} R) := AB4.of_AB5 _ + +instance : AB4Star (ModuleCat.{u} R) where + ofShape J := + HasExactLimitsOfShape.domain_of_functor (Discrete J) (forget₂ (ModuleCat R) AddCommGrp.{u}) From b0a22ed4c5641d49feb0bd02f469d2f0e0cc5daf Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 13 Dec 2024 16:45:31 +0000 Subject: [PATCH 702/829] fix: ensure the clashing `X[i]` notations have the same precedence (#19900) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit I don't exactly know why core has `:max` for `GetElem`, but not doing so here makes the errors more confusing when both notations are available. This is needed to make ```lean import Mathlib open scoped ProbabilityTheory example (l : List (Fin 3 → ℕ)) (i) (hi : i < l.length) (j : Fin 3) : 0 ≤ l[i] j := sorry ``` be legal. This also cleans up some porting notes which seem to no longer apply, and adds `noWs` to make the notation ambiguity a little less confusing. --- Mathlib/Probability/Notation.lean | 12 ++++++++--- Mathlib/Probability/StrongLaw.lean | 34 +++++++++++++++--------------- Mathlib/Probability/Variance.lean | 8 ++----- 3 files changed, 28 insertions(+), 26 deletions(-) diff --git a/Mathlib/Probability/Notation.lean b/Mathlib/Probability/Notation.lean index dc866c19c326d..aaf254c1f9096 100644 --- a/Mathlib/Probability/Notation.lean +++ b/Mathlib/Probability/Notation.lean @@ -38,9 +38,15 @@ open scoped MeasureTheory scoped[ProbabilityTheory] notation "𝔼[" X "|" m "]" => MeasureTheory.condexp m MeasureTheory.MeasureSpace.volume X --- Note(kmill): this notation tends to lead to ambiguity with GetElem notation. -set_option quotPrecheck false in -scoped[ProbabilityTheory] notation P "[" X "]" => ∫ x, ↑(X x) ∂P +-- `scoped[ProbabilityTheory]` isn't legal for `macro`s. +namespace ProbabilityTheory +/-- `P[X]` is the expectation of `X` under the measure `P`. + +Note that this notation can conflict with the `GetElem` notation for lists. Usually if you see an +error about ambiguous notation when trying to write `l[i]` for a list, it means that Lean could +not find `i < l.length`, and so fell back to trying this notation as well. -/ +scoped macro:max P:term noWs "[" X:term "]" : term => `(∫ x, ↑($X x) ∂$P) +end ProbabilityTheory scoped[ProbabilityTheory] notation "𝔼[" X "]" => ∫ a, (X : _ → _) a diff --git a/Mathlib/Probability/StrongLaw.lean b/Mathlib/Probability/StrongLaw.lean index 8c69211cebfef..734ac6c8622b2 100644 --- a/Mathlib/Probability/StrongLaw.lean +++ b/Mathlib/Probability/StrongLaw.lean @@ -606,7 +606,7 @@ theorem strong_law_ae_real {Ω : Type*} {m : MeasurableSpace Ω} {μ : Measure (X : ℕ → Ω → ℝ) (hint : Integrable (X 0) μ) (hindep : Pairwise ((IndepFun · · μ) on X)) (hident : ∀ i, IdentDistrib (X i) (X 0) μ μ) : - ∀ᵐ ω ∂μ, Tendsto (fun n : ℕ => (∑ i ∈ range n, X i ω) / n) atTop (𝓝 (μ [X 0])) := by + ∀ᵐ ω ∂μ, Tendsto (fun n : ℕ => (∑ i ∈ range n, X i ω) / n) atTop (𝓝 μ[X 0]) := by let mΩ : MeasureSpace Ω := ⟨μ⟩ -- first get rid of the trivial case where the space is not a probability space by_cases h : ∀ᵐ ω, X 0 ω = 0 @@ -654,12 +654,12 @@ lemma strong_law_ae_simpleFunc_comp (X : ℕ → Ω → E) (h' : Measurable (X 0 (hindep : Pairwise ((IndepFun · · μ) on X)) (hident : ∀ i, IdentDistrib (X i) (X 0) μ μ) (φ : SimpleFunc E E) : ∀ᵐ ω ∂μ, - Tendsto (fun n : ℕ ↦ (n : ℝ) ⁻¹ • (∑ i ∈ range n, φ (X i ω))) atTop (𝓝 (μ [φ ∘ (X 0)])) := by + Tendsto (fun n : ℕ ↦ (n : ℝ) ⁻¹ • (∑ i ∈ range n, φ (X i ω))) atTop (𝓝 μ[φ ∘ (X 0)]) := by -- this follows from the one-dimensional version when `φ` takes a single value, and is then -- extended to the general case by linearity. classical refine SimpleFunc.induction (P := fun ψ ↦ ∀ᵐ ω ∂μ, - Tendsto (fun n : ℕ ↦ (n : ℝ) ⁻¹ • (∑ i ∈ range n, ψ (X i ω))) atTop (𝓝 (μ [ψ ∘ (X 0)]))) ?_ ?_ φ + Tendsto (fun n : ℕ ↦ (n : ℝ) ⁻¹ • (∑ i ∈ range n, ψ (X i ω))) atTop (𝓝 μ[ψ ∘ (X 0)])) ?_ ?_ φ · intro c s hs simp only [SimpleFunc.const_zero, SimpleFunc.coe_piecewise, SimpleFunc.coe_const, SimpleFunc.coe_zero, piecewise_eq_indicator, Function.comp_apply] @@ -667,7 +667,7 @@ lemma strong_law_ae_simpleFunc_comp (X : ℕ → Ω → E) (h' : Measurable (X 0 have F_meas : Measurable F := (measurable_indicator_const_iff 1).2 hs let Y : ℕ → Ω → ℝ := fun n ↦ F ∘ (X n) have : ∀ᵐ (ω : Ω) ∂μ, Tendsto (fun (n : ℕ) ↦ (n : ℝ)⁻¹ • ∑ i ∈ Finset.range n, Y i ω) - atTop (𝓝 (μ [Y 0])) := by + atTop (𝓝 μ[Y 0]) := by simp only [Function.const_one, smul_eq_mul, ← div_eq_inv_mul] apply strong_law_ae_real · exact SimpleFunc.integrable_of_isFiniteMeasure @@ -704,7 +704,7 @@ lemma strong_law_ae_of_measurable (X : ℕ → Ω → E) (hint : Integrable (X 0) μ) (h' : StronglyMeasurable (X 0)) (hindep : Pairwise ((IndepFun · · μ) on X)) (hident : ∀ i, IdentDistrib (X i) (X 0) μ μ) : - ∀ᵐ ω ∂μ, Tendsto (fun n : ℕ ↦ (n : ℝ) ⁻¹ • (∑ i ∈ range n, X i ω)) atTop (𝓝 (μ [X 0])) := by + ∀ᵐ ω ∂μ, Tendsto (fun n : ℕ ↦ (n : ℝ) ⁻¹ • (∑ i ∈ range n, X i ω)) atTop (𝓝 μ[X 0]) := by /- Choose a simple function `φ` such that `φ (X 0)` approximates well enough `X 0` -- this is possible as `X 0` is strongly measurable. Then `φ (X n)` approximates well `X n`. Then the strong law for `φ (X n)` implies the strong law for `X n`, up to a small @@ -721,11 +721,11 @@ lemma strong_law_ae_of_measurable let Y : ℕ → ℕ → Ω → E := fun k i ↦ (φ k) ∘ (X i) -- strong law for `φ (X n)` have A : ∀ᵐ ω ∂μ, ∀ k, - Tendsto (fun n : ℕ ↦ (n : ℝ) ⁻¹ • (∑ i ∈ range n, Y k i ω)) atTop (𝓝 (μ [Y k 0])) := + Tendsto (fun n : ℕ ↦ (n : ℝ) ⁻¹ • (∑ i ∈ range n, Y k i ω)) atTop (𝓝 μ[Y k 0]) := ae_all_iff.2 (fun k ↦ strong_law_ae_simpleFunc_comp X h'.measurable hindep hident (φ k)) -- strong law for the error `‖X i - φ (X i)‖` have B : ∀ᵐ ω ∂μ, ∀ k, Tendsto (fun n : ℕ ↦ (∑ i ∈ range n, ‖(X i - Y k i) ω‖) / n) - atTop (𝓝 (μ [(fun ω ↦ ‖(X 0 - Y k 0) ω‖)])) := by + atTop (𝓝 μ[(fun ω ↦ ‖(X 0 - Y k 0) ω‖)]) := by apply ae_all_iff.2 (fun k ↦ ?_) let G : ℕ → E → ℝ := fun k x ↦ ‖x - φ k x‖ have G_meas : ∀ k, Measurable (G k) := @@ -753,28 +753,28 @@ lemma strong_law_ae_of_measurable simp_rw [Pi.sub_apply, norm_sub_rev (X 0 _)] exact ((tendsto_order.1 (tendsto_integral_norm_approxOn_sub h'.measurable hint)).2 δ δpos).exists - have : ‖μ [Y k 0] - μ [X 0]‖ < δ := by + have : ‖μ[Y k 0] - μ[X 0]‖ < δ := by rw [norm_sub_rev, ← integral_sub hint] · exact (norm_integral_le_integral_norm _).trans_lt hk · exact ((φ k).comp (X 0) h'.measurable).integrable_of_isFiniteMeasure -- consider `n` large enough for which the above convergences have taken place within `δ`. have I : ∀ᶠ n in atTop, (∑ i ∈ range n, ‖(X i - Y k i) ω‖) / n < δ := (tendsto_order.1 (h'ω k)).2 δ hk - have J : ∀ᶠ (n : ℕ) in atTop, ‖(n : ℝ) ⁻¹ • (∑ i ∈ range n, Y k i ω) - μ [Y k 0]‖ < δ := by + have J : ∀ᶠ (n : ℕ) in atTop, ‖(n : ℝ) ⁻¹ • (∑ i ∈ range n, Y k i ω) - μ[Y k 0]‖ < δ := by specialize hω k rw [tendsto_iff_norm_sub_tendsto_zero] at hω exact (tendsto_order.1 hω).2 δ δpos filter_upwards [I, J] with n hn h'n -- at such an `n`, the strong law is realized up to `ε`. calc - ‖(n : ℝ)⁻¹ • ∑ i ∈ Finset.range n, X i ω - μ [X 0]‖ + ‖(n : ℝ)⁻¹ • ∑ i ∈ Finset.range n, X i ω - μ[X 0]‖ = ‖(n : ℝ)⁻¹ • ∑ i ∈ Finset.range n, (X i ω - Y k i ω) + - ((n : ℝ)⁻¹ • ∑ i ∈ Finset.range n, Y k i ω - μ [Y k 0]) + (μ [Y k 0] - μ [X 0])‖ := by + ((n : ℝ)⁻¹ • ∑ i ∈ Finset.range n, Y k i ω - μ[Y k 0]) + (μ[Y k 0] - μ[X 0])‖ := by congr simp only [Function.comp_apply, sum_sub_distrib, smul_sub] abel _ ≤ ‖(n : ℝ)⁻¹ • ∑ i ∈ Finset.range n, (X i ω - Y k i ω)‖ + - ‖(n : ℝ)⁻¹ • ∑ i ∈ Finset.range n, Y k i ω - μ [Y k 0]‖ + ‖μ [Y k 0] - μ [X 0]‖ := + ‖(n : ℝ)⁻¹ • ∑ i ∈ Finset.range n, Y k i ω - μ[Y k 0]‖ + ‖μ[Y k 0] - μ[X 0]‖ := norm_add₃_le _ ≤ (∑ i ∈ Finset.range n, ‖X i ω - Y k i ω‖) / n + δ + δ := by gcongr @@ -795,7 +795,7 @@ version, due to Etemadi, that only requires pairwise independence. -/ theorem strong_law_ae (X : ℕ → Ω → E) (hint : Integrable (X 0) μ) (hindep : Pairwise ((IndepFun · · μ) on X)) (hident : ∀ i, IdentDistrib (X i) (X 0) μ μ) : - ∀ᵐ ω ∂μ, Tendsto (fun n : ℕ ↦ (n : ℝ) ⁻¹ • (∑ i ∈ range n, X i ω)) atTop (𝓝 (μ [X 0])) := by + ∀ᵐ ω ∂μ, Tendsto (fun n : ℕ ↦ (n : ℝ) ⁻¹ • (∑ i ∈ range n, X i ω)) atTop (𝓝 μ[X 0]) := by -- First exclude the trivial case where the space is not a probability space by_cases h : ∀ᵐ ω ∂μ, X 0 ω = 0 · have I : ∀ᵐ ω ∂μ, ∀ i, X i ω = 0 := by @@ -814,12 +814,12 @@ theorem strong_law_ae (X : ℕ → Ω → E) (hint : Integrable (X 0) μ) ae_all_iff.2 (fun i ↦ AEStronglyMeasurable.ae_eq_mk (A i).1) have Yint : Integrable (Y 0) μ := Integrable.congr hint (AEStronglyMeasurable.ae_eq_mk (A 0).1) have C : ∀ᵐ ω ∂μ, - Tendsto (fun n : ℕ ↦ (n : ℝ) ⁻¹ • (∑ i ∈ range n, Y i ω)) atTop (𝓝 (μ [Y 0])) := by + Tendsto (fun n : ℕ ↦ (n : ℝ) ⁻¹ • (∑ i ∈ range n, Y i ω)) atTop (𝓝 μ[Y 0]) := by apply strong_law_ae_of_measurable Y Yint ((A 0).1.stronglyMeasurable_mk) (fun i j hij ↦ IndepFun.ae_eq (hindep hij) (A i).1.ae_eq_mk (A j).1.ae_eq_mk) (fun i ↦ ((A i).1.identDistrib_mk.symm.trans (hident i)).trans (A 0).1.identDistrib_mk) filter_upwards [B, C] with ω h₁ h₂ - have : μ [X 0] = μ [Y 0] := integral_congr_ae (AEStronglyMeasurable.ae_eq_mk (A 0).1) + have : μ[X 0] = μ[Y 0] := integral_congr_ae (AEStronglyMeasurable.ae_eq_mk (A 0).1) rw [this] apply Tendsto.congr (fun n ↦ ?_) h₂ congr with i @@ -839,7 +839,7 @@ converges in `Lᵖ` to `𝔼[X 0]`. -/ theorem strong_law_Lp {p : ℝ≥0∞} (hp : 1 ≤ p) (hp' : p ≠ ∞) (X : ℕ → Ω → E) (hℒp : Memℒp (X 0) p μ) (hindep : Pairwise ((IndepFun · · μ) on X)) (hident : ∀ i, IdentDistrib (X i) (X 0) μ μ) : - Tendsto (fun (n : ℕ) => eLpNorm (fun ω => (n : ℝ) ⁻¹ • (∑ i ∈ range n, X i ω) - μ [X 0]) p μ) + Tendsto (fun (n : ℕ) => eLpNorm (fun ω => (n : ℝ) ⁻¹ • (∑ i ∈ range n, X i ω) - μ[X 0]) p μ) atTop (𝓝 0) := by -- First exclude the trivial case where the space is not a probability space by_cases h : ∀ᵐ ω ∂μ, X 0 ω = 0 @@ -847,7 +847,7 @@ theorem strong_law_Lp {p : ℝ≥0∞} (hp : 1 ≤ p) (hp' : p ≠ ∞) (X : ℕ rw [ae_all_iff] intro i exact (hident i).symm.ae_snd (p := fun x ↦ x = 0) measurableSet_eq h - have A (n : ℕ) : eLpNorm (fun ω => (n : ℝ) ⁻¹ • (∑ i ∈ range n, X i ω) - μ [X 0]) p μ = 0 := by + have A (n : ℕ) : eLpNorm (fun ω => (n : ℝ) ⁻¹ • (∑ i ∈ range n, X i ω) - μ[X 0]) p μ = 0 := by simp only [integral_eq_zero_of_ae h, sub_zero] apply eLpNorm_eq_zero_of_ae_zero filter_upwards [I] with ω hω diff --git a/Mathlib/Probability/Variance.lean b/Mathlib/Probability/Variance.lean index 00263fc96ac34..b5cc11cbd0014 100644 --- a/Mathlib/Probability/Variance.lean +++ b/Mathlib/Probability/Variance.lean @@ -77,9 +77,7 @@ theorem evariance_eq_top [IsFiniteMeasure μ] (hXm : AEStronglyMeasurable X μ) simp only [ENNReal.toReal_ofNat, ENNReal.one_toReal, ENNReal.rpow_two, Ne] exact ENNReal.rpow_lt_top_of_nonneg (by linarith) h.ne refine hX ?_ - -- Porting note: `μ[X]` without whitespace is ambiguous as it could be GetElem, - -- and `convert` cannot disambiguate based on typeclass inference failure. - convert this.add (memℒp_const <| μ [X]) + convert this.add (memℒp_const μ[X]) ext ω rw [Pi.add_apply, sub_add_cancel] @@ -120,9 +118,7 @@ theorem _root_.MeasureTheory.Memℒp.variance_eq [IsFiniteMeasure μ] (hX : Mem rw [variance, evariance_eq_lintegral_ofReal, ← ofReal_integral_eq_lintegral_ofReal, ENNReal.toReal_ofReal (by positivity)] · rfl - · -- Porting note: `μ[X]` without whitespace is ambiguous as it could be GetElem, - -- and `convert` cannot disambiguate based on typeclass inference failure. - convert (hX.sub <| memℒp_const (μ [X])).integrable_norm_rpow two_ne_zero ENNReal.two_ne_top + · convert (hX.sub <| memℒp_const μ[X]).integrable_norm_rpow two_ne_zero ENNReal.two_ne_top with ω simp only [Pi.sub_apply, Real.norm_eq_abs, ENNReal.toReal_ofNat, ENNReal.one_toReal, Real.rpow_two, sq_abs, abs_pow] From 5aaaf8b3787cb2bdbe430891f67b07e9c3bf24ed Mon Sep 17 00:00:00 2001 From: Nick Ward <102917377+gio256@users.noreply.github.com> Date: Fri, 13 Dec 2024 19:15:06 +0000 Subject: [PATCH 703/829] feat(CategoryTheory/Enriched): Add eHomCongr and related lemmas (#19932) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Recall that when `C` is both a category and a `V`-enriched category, we say it is an `EnrichedOrdinaryCategory` if it comes equipped with an equivalence between morphisms `X ⟶ Y` in `C` and morphisms `𝟙_ V ⟶ (X ⟶[V] Y)` in `V` that preserves identities and composition in `C`. In such a `V`-enriched ordinary category `C`, isomorphisms in `C` induce isomorphisms between hom-objects in `V`. We define this isomorphism, `CategoryTheory.Iso.eHomCongr`, and prove that it behaves nicely with respect to composition in `C`. The lemmas here parallel those for unenriched categories in [`Mathlib/CategoryTheory/HomCongr.lean`](https://github.com/leanprover-community/mathlib4/blob/a1c80a59b69da977f8519230ef96b3b95e31bc1d/Mathlib/CategoryTheory/HomCongr.lean) and those for sorts in [`Mathlib/Logic/Equiv/Defs.lean`](https://github.com/leanprover-community/mathlib4/blob/a1c80a59b69da977f8519230ef96b3b95e31bc1d/Mathlib/Logic/Equiv/Defs.lean#L428) (cf. `Equiv.arrowCongr`). Note, however, that they construct equivalences between `Type`s and `Sort`s, respectively, while in this PR we construct isomorphisms between objects in `V`. --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Enriched/HomCongr.lean | 95 +++++++++++++++++++ Mathlib/CategoryTheory/Enriched/Ordinary.lean | 38 ++++++++ 3 files changed, 134 insertions(+) create mode 100644 Mathlib/CategoryTheory/Enriched/HomCongr.lean diff --git a/Mathlib.lean b/Mathlib.lean index dfa70e2924e2d..f26acaca225ac 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1706,6 +1706,7 @@ import Mathlib.CategoryTheory.Endofunctor.Algebra import Mathlib.CategoryTheory.Endomorphism import Mathlib.CategoryTheory.Enriched.Basic import Mathlib.CategoryTheory.Enriched.FunctorCategory +import Mathlib.CategoryTheory.Enriched.HomCongr import Mathlib.CategoryTheory.Enriched.Opposite import Mathlib.CategoryTheory.Enriched.Ordinary import Mathlib.CategoryTheory.EpiMono diff --git a/Mathlib/CategoryTheory/Enriched/HomCongr.lean b/Mathlib/CategoryTheory/Enriched/HomCongr.lean new file mode 100644 index 0000000000000..1d4ab2b6588ea --- /dev/null +++ b/Mathlib/CategoryTheory/Enriched/HomCongr.lean @@ -0,0 +1,95 @@ +/- +Copyright (c) 2024 Nick Ward. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Nick Ward +-/ +import Mathlib.CategoryTheory.Enriched.Ordinary + +/-! +# Congruence of enriched homs + +Recall that when `C` is both a category and a `V`-enriched category, we say it +is an `EnrichedOrdinaryCategory` if it comes equipped with a sufficiently +compatible equivalence between morphisms `X ⟶ Y` in `C` and morphisms +`𝟙_ V ⟶ (X ⟶[V] Y)` in `V`. + +In such a `V`-enriched ordinary category `C`, isomorphisms in `C` induce +isomorphisms between hom-objects in `V`. We define this isomorphism in +`CategoryTheory.Iso.eHomCongr` and prove that it respects composition in `C`. + +The treatment here parallels that for unenriched categories in +`Mathlib/CategoryTheory/HomCongr.lean` and that for sorts in +`Mathlib/Logic/Equiv/Defs.lean` (cf. `Equiv.arrowCongr`). Note, however, that +they construct equivalences between `Type`s and `Sort`s, respectively, while +in this file we construct isomorphisms between objects in `V`. +-/ + +universe v' v u u' + +namespace CategoryTheory +namespace Iso + +open Category MonoidalCategory + +variable (V : Type u') [Category.{v'} V] [MonoidalCategory V] + {C : Type u} [Category.{v} C] [EnrichedOrdinaryCategory V C] + +/-- Given isomorphisms `α : X ≅ X₁` and `β : Y ≅ Y₁` in `C`, we can construct +an isomorphism between `V` objects `X ⟶[V] Y` and `X₁ ⟶[V] Y₁`. -/ +@[simps] +def eHomCongr {X Y X₁ Y₁ : C} (α : X ≅ X₁) (β : Y ≅ Y₁) : + (X ⟶[V] Y) ≅ (X₁ ⟶[V] Y₁) where + hom := eHomWhiskerRight V α.inv Y ≫ eHomWhiskerLeft V X₁ β.hom + inv := eHomWhiskerRight V α.hom Y₁ ≫ eHomWhiskerLeft V X β.inv + hom_inv_id := by + rw [← eHom_whisker_exchange] + slice_lhs 2 3 => rw [← eHomWhiskerRight_comp] + simp [← eHomWhiskerLeft_comp] + inv_hom_id := by + rw [← eHom_whisker_exchange] + slice_lhs 2 3 => rw [← eHomWhiskerRight_comp] + simp [← eHomWhiskerLeft_comp] + +lemma eHomCongr_refl (X Y : C) : + eHomCongr V (Iso.refl X) (Iso.refl Y) = Iso.refl (X ⟶[V] Y) := by aesop + +lemma eHomCongr_trans {X₁ Y₁ X₂ Y₂ X₃ Y₃ : C} (α₁ : X₁ ≅ X₂) (β₁ : Y₁ ≅ Y₂) + (α₂ : X₂ ≅ X₃) (β₂ : Y₂ ≅ Y₃) : + eHomCongr V (α₁ ≪≫ α₂) (β₁ ≪≫ β₂) = + eHomCongr V α₁ β₁ ≪≫ eHomCongr V α₂ β₂ := by + ext; simp [eHom_whisker_exchange_assoc] + +lemma eHomCongr_symm {X Y X₁ Y₁ : C} (α : X ≅ X₁) (β : Y ≅ Y₁) : + (eHomCongr V α β).symm = eHomCongr V α.symm β.symm := rfl + +/-- `eHomCongr` respects composition of morphisms. Recall that for any +composable pair of arrows `f : X ⟶ Y` and `g : Y ⟶ Z` in `C`, the composite +`f ≫ g` in `C` defines a morphism `𝟙_ V ⟶ (X ⟶[V] Z)` in `V`. Composing with +the isomorphism `eHomCongr V α γ` yields a morphism in `V` that can be factored +through the enriched composition map as shown: +`𝟙_ V ⟶ 𝟙_ V ⊗ 𝟙_ V ⟶ (X₁ ⟶[V] Y₁) ⊗ (Y₁ ⟶[V] Z₁) ⟶ (X₁ ⟶[V] Z₁)`. -/ +@[reassoc] +lemma eHomCongr_comp {X Y Z X₁ Y₁ Z₁ : C} (α : X ≅ X₁) (β : Y ≅ Y₁) (γ : Z ≅ Z₁) + (f : X ⟶ Y) (g : Y ⟶ Z) : + eHomEquiv V (f ≫ g) ≫ (eHomCongr V α γ).hom = + (λ_ _).inv ≫ (eHomEquiv V f ≫ (eHomCongr V α β).hom) ▷ _ ≫ + _ ◁ (eHomEquiv V g ≫ (eHomCongr V β γ).hom) ≫ eComp V X₁ Y₁ Z₁ := by + simp only [eHomCongr, MonoidalCategory.whiskerRight_id, assoc, + MonoidalCategory.whiskerLeft_comp] + rw [rightUnitor_inv_naturality_assoc, rightUnitor_inv_naturality_assoc, + rightUnitor_inv_naturality_assoc, hom_inv_id_assoc, ← whisker_exchange_assoc, + ← whisker_exchange_assoc, ← eComp_eHomWhiskerLeft, eHom_whisker_cancel_assoc, + ← eComp_eHomWhiskerRight_assoc, ← tensorHom_def_assoc, + ← eHomEquiv_comp_assoc] + +/-- The inverse map defined by `eHomCongr` respects composition of morphisms. -/ +@[reassoc] +lemma eHomCongr_inv_comp {X Y Z X₁ Y₁ Z₁ : C} (α : X ≅ X₁) (β : Y ≅ Y₁) + (γ : Z ≅ Z₁) (f : X₁ ⟶ Y₁) (g : Y₁ ⟶ Z₁) : + eHomEquiv V (f ≫ g) ≫ (eHomCongr V α γ).inv = + (λ_ _).inv ≫ (eHomEquiv V f ≫ (eHomCongr V α β).inv) ▷ _ ≫ + _ ◁ (eHomEquiv V g ≫ (eHomCongr V β γ).inv) ≫ eComp V X Y Z := + eHomCongr_comp V α.symm β.symm γ.symm f g + +end Iso +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Enriched/Ordinary.lean b/Mathlib/CategoryTheory/Enriched/Ordinary.lean index 85c4db8068474..031e85055e4ac 100644 --- a/Mathlib/CategoryTheory/Enriched/Ordinary.lean +++ b/Mathlib/CategoryTheory/Enriched/Ordinary.lean @@ -81,6 +81,15 @@ lemma eHomWhiskerRight_comp {X X' X'' : C} (f : X ⟶ X') (f' : X' ⟶ X'') (Y : ← associator_inv_naturality_left_assoc, Iso.inv_hom_id_assoc, ← whisker_exchange_assoc, id_whiskerLeft_assoc, Iso.inv_hom_id_assoc] +/-- Whiskering commutes with the enriched composition. -/ +@[reassoc] +lemma eComp_eHomWhiskerRight {X X' : C} (f : X ⟶ X') (Y Z : C) : + eComp V X' Y Z ≫ eHomWhiskerRight V f Z = + eHomWhiskerRight V f Y ▷ _ ≫ eComp V X Y Z := by + dsimp [eHomWhiskerRight] + rw [leftUnitor_inv_naturality_assoc, whisker_exchange_assoc] + simp [e_assoc'] + /-- The morphism `(X ⟶[V] Y) ⟶ (X ⟶[V] Y')` induced by a morphism `Y ⟶ Y'`. -/ def eHomWhiskerLeft (X : C) {Y Y' : C} (g : Y ⟶ Y') : (X ⟶[V] Y) ⟶ (X ⟶[V] Y') := @@ -104,6 +113,35 @@ lemma eHomWhiskerLeft_comp (X : C) {Y Y' Y'' : C} (g : Y ⟶ Y') (g' : Y' ⟶ Y' associator_inv_naturality_right_assoc, Iso.hom_inv_id_assoc, whisker_exchange_assoc, MonoidalCategory.whiskerRight_id_assoc, Iso.inv_hom_id_assoc] +/-- Whiskering commutes with the enriched composition. -/ +@[reassoc] +lemma eComp_eHomWhiskerLeft (X Y : C) {Z Z' : C} (g : Z ⟶ Z') : + eComp V X Y Z ≫ eHomWhiskerLeft V X g = + _ ◁ eHomWhiskerLeft V Y g ≫ eComp V X Y Z' := by + dsimp [eHomWhiskerLeft] + rw [rightUnitor_inv_naturality_assoc, ← whisker_exchange_assoc] + simp [e_assoc'] + +/-- Given an isomorphism `α : Y ≅ Y₁` in C, the enriched composition map +`eComp V X Y Z : (X ⟶[V] Y) ⊗ (Y ⟶[V] Z) ⟶ (X ⟶[V] Z)` factors through the `V` +object `(X ⟶[V] Y₁) ⊗ (Y₁ ⟶[V] Z)` via the map defined by whiskering in the +middle with `α.hom` and `α.inv`. -/ +@[reassoc] +lemma eHom_whisker_cancel {X Y Y₁ Z : C} (α : Y ≅ Y₁) : + eHomWhiskerLeft V X α.hom ▷ _ ≫ _ ◁ eHomWhiskerRight V α.inv Z ≫ + eComp V X Y₁ Z = eComp V X Y Z := by + dsimp [eHomWhiskerLeft, eHomWhiskerRight] + simp only [MonoidalCategory.whiskerLeft_comp_assoc, whisker_assoc_symm, + triangle_assoc_comp_left_inv_assoc, e_assoc', assoc] + simp only [← comp_whiskerRight_assoc] + change (eHomWhiskerLeft V X α.hom ≫ eHomWhiskerLeft V X α.inv) ▷ _ ≫ _ = _ + simp [← eHomWhiskerLeft_comp] + +@[reassoc] +lemma eHom_whisker_cancel_inv {X Y Y₁ Z : C} (α : Y ≅ Y₁) : + eHomWhiskerLeft V X α.inv ▷ _ ≫ _ ◁ eHomWhiskerRight V α.hom Z ≫ + eComp V X Y Z = eComp V X Y₁ Z := eHom_whisker_cancel V α.symm + @[reassoc] lemma eHom_whisker_exchange {X X' Y Y' : C} (f : X ⟶ X') (g : Y ⟶ Y') : eHomWhiskerLeft V X' g ≫ eHomWhiskerRight V f Y' = From 76cecbbdcd76496b9cb042e9e11252b18a845566 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Fri, 13 Dec 2024 19:24:18 +0000 Subject: [PATCH 704/829] feat(Topology/Algebra/Group/Basic): a quotient map between topological groups is open (#19657) From the FLT project. --- Mathlib/Topology/Algebra/Group/Basic.lean | 33 +++++++++++++++++++++++ 1 file changed, 33 insertions(+) diff --git a/Mathlib/Topology/Algebra/Group/Basic.lean b/Mathlib/Topology/Algebra/Group/Basic.lean index 6746a7d253f9f..e37378c9fa050 100644 --- a/Mathlib/Topology/Algebra/Group/Basic.lean +++ b/Mathlib/Topology/Algebra/Group/Basic.lean @@ -771,6 +771,39 @@ theorem continuous_of_continuousAt_one₂ {H M : Type*} [CommMonoid M] [Topologi (((hl x).comp tendsto_snd).mul hf)).mono_right (le_of_eq ?_) simp only [map_one, mul_one, MonoidHom.one_apply] +-- TODO: unify with `QuotientGroup.isOpenQuotientMap_mk` +/-- Let `A` and `B` be topological groups, and let `φ : A → B` be a continuous surjective group +homomorphism. Assume furthermore that `φ` is a quotient map (i.e., `V ⊆ B` +is open iff `φ⁻¹ V` is open). Then `φ` is an open quotient map, and in particular an open map. -/ +@[to_additive "Let `A` and `B` be topological additive groups, and let `φ : A → B` be a continuous +surjective additive group homomorphism. Assume furthermore that `φ` is a quotient map (i.e., `V ⊆ B` +is open iff `φ⁻¹ V` is open). Then `φ` is an open quotient map, and in particular an open map."] +lemma MonoidHom.isOpenQuotientMap_of_isQuotientMap {A : Type*} [Group A] + [TopologicalSpace A] [TopologicalGroup A] {B : Type*} [Group B] [TopologicalSpace B] + {F : Type*} [FunLike F A B] [MonoidHomClass F A B] {φ : F} + (hφ : IsQuotientMap φ) : IsOpenQuotientMap φ where + surjective := hφ.surjective + continuous := hφ.continuous + isOpenMap := by + -- We need to check that if `U ⊆ A` is open then `φ⁻¹ (φ U)` is open. + intro U hU + rw [← hφ.isOpen_preimage] + -- It suffices to show that `φ⁻¹ (φ U) = ⋃ (U * k⁻¹)` as `k` runs through the kernel of `φ`, + -- as `U * k⁻¹` is open because `x ↦ x * k` is continuous. + -- Remark: here is where we use that we have groups not monoids (you cannot avoid + -- using both `k` and `k⁻¹` at this point). + suffices ⇑φ ⁻¹' (⇑φ '' U) = ⋃ k ∈ ker (φ : A →* B), (fun x ↦ x * k) ⁻¹' U by + exact this ▸ isOpen_biUnion (fun k _ ↦ Continuous.isOpen_preimage (by fun_prop) _ hU) + ext x + -- But this is an elementary calculation. + constructor + · rintro ⟨y, hyU, hyx⟩ + apply Set.mem_iUnion_of_mem (x⁻¹ * y) + simp_all + · rintro ⟨_, ⟨k, rfl⟩, _, ⟨(hk : φ k = 1), rfl⟩, hx⟩ + use x * k, hx + rw [map_mul, hk, mul_one] + @[to_additive] theorem TopologicalGroup.ext {G : Type*} [Group G] {t t' : TopologicalSpace G} (tg : @TopologicalGroup G t _) (tg' : @TopologicalGroup G t' _) From 73ea8d5d58fb932e9b4746030d794df128750427 Mon Sep 17 00:00:00 2001 From: "Filippo A. E. Nuccio" Date: Fri, 13 Dec 2024 19:45:44 +0000 Subject: [PATCH 705/829] feat(NumberTheory/Padics/RingHoms): add the equivalence between the residue field of Z_p and ZMod p (#19922) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We construct the equivalence between the residue field of Z_p and ZMod p in terms of the `IsLocalRing.ResidueField` declaration. Co-authored-by: María Inés de Frutos Fernández @mariainesdff --- Mathlib/NumberTheory/Padics/RingHoms.lean | 29 ++++++++++++++--------- 1 file changed, 18 insertions(+), 11 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/RingHoms.lean b/Mathlib/NumberTheory/Padics/RingHoms.lean index 53244367555c2..106a4b2f15a8e 100644 --- a/Mathlib/NumberTheory/Padics/RingHoms.lean +++ b/Mathlib/NumberTheory/Padics/RingHoms.lean @@ -3,36 +3,38 @@ Copyright (c) 2020 Johan Commelin, Robert Y. Lewis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Robert Y. Lewis -/ +import Mathlib.RingTheory.LocalRing.ResidueField.Defs import Mathlib.NumberTheory.Padics.PadicIntegers import Mathlib.RingTheory.ZMod /-! -# Relating `ℤ_[p]` to `ZMod (p ^ n)` +# Relating `ℤ_[p]` to `ZMod (p ^ n)`, aka `ℤ/p^nℤ`. -In this file we establish connections between the `p`-adic integers $\mathbb{Z}_p$ -and the integers modulo powers of `p`, $\mathbb{Z}/p^n\mathbb{Z}$. +In this file we establish connections between the `p`-adic integers `ℤ_[p]` +and the integers modulo powers of `p`, `ℤ/p^nℤ`, implemented as `ZMod (p^n)`. ## Main declarations -We show that $\mathbb{Z}_p$ has a ring hom to $\mathbb{Z}/p^n\mathbb{Z}$ for each `n`. +We show that `ℤ_[p]` has a ring homomorphism to `ℤ/p^nℤ` for each `n`. The case for `n = 1` is handled separately, since it is used in the general construction and we may want to use it without the `^1` getting in the way. -* `PadicInt.toZMod`: ring hom to `ZMod p` -* `PadicInt.toZModPow`: ring hom to `ZMod (p^n)` +* `PadicInt.toZMod`: ring homomorphism to `ℤ/pℤ`, implemented as `ZMod p`. +* `PadicInt.toZModPow`: ring homomorphism to `ℤ/p^nℤ`, implemented as `ZMod (p^n)`. * `PadicInt.ker_toZMod` / `PadicInt.ker_toZModPow`: the kernels of these maps are the ideals generated by `p^n` +* `PadicInt.residueField` shows that the residue field of `ℤ_[p]` is isomorhic to ``ℤ/pℤ`. -We also establish the universal property of $\mathbb{Z}_p$ as a projective limit. -Given a family of compatible ring homs $f_k : R \to \mathbb{Z}/p^n\mathbb{Z}$, -there is a unique limit $R \to \mathbb{Z}_p$. +We also establish the universal property of `ℤ_[p]` as a projective limit. +Given a family of compatible ring homomorphisms `f_k : R → ℤ/p^nℤ`, +there is a unique limit `R → ℤ_[p]` * `PadicInt.lift`: the limit function * `PadicInt.lift_spec` / `PadicInt.lift_unique`: the universal property ## Implementation notes -The ring hom constructions go through an auxiliary constructor `PadicInt.toZModHom`, -which removes some boilerplate code. +The constructions of the ring homomorphisms go through an auxiliary constructor +`PadicInt.toZModHom`, which removes some boilerplate code. -/ @@ -284,6 +286,11 @@ theorem ker_toZMod : RingHom.ker (toZMod : ℤ_[p] →+* ZMod p) = maximalIdeal · norm_cast · apply sub_zmodRepr_mem +/-- The equivalence between the residue field of the `p`-adic integers and `ℤ/pℤ` -/ +def residueField : IsLocalRing.ResidueField ℤ_[p] ≃+* ZMod p := by + exact_mod_cast (@PadicInt.ker_toZMod p _) ▸ RingHom.quotientKerEquivOfSurjective + (ZMod.ringHom_surjective PadicInt.toZMod) + /-- `appr n x` gives a value `v : ℕ` such that `x` and `↑v : ℤ_p` are congruent mod `p^n`. See `appr_spec`. -/ -- Porting note: removing irreducible solves a lot of problems From d1fa45bcb0bb5ae74fc38931a796fba555e94e41 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Fri, 13 Dec 2024 20:24:57 +0000 Subject: [PATCH 706/829] =?UTF-8?q?feat(FieldTheory/Finite/Basic):=20`(1?= =?UTF-8?q?=20+=20p=20*=20a)=20^=20(p=20^=20m)=20=E2=89=A1=201=20[MOD=20p?= =?UTF-8?q?=20^=20m]`=20(#19881)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds `(1 + p * a) ^ (p ^ m) ≡ 1 [MOD p ^ m]` and deduces that if `n` is a prime power and `a : ZMod n` has order prime to `n`, then either `a = 1` or `a - 1` is a unit. This will be used in the proof that the commutator subgroup of a Z-group is a Hall subgroup. Co-authored-by: Thomas Browning --- Mathlib/FieldTheory/Finite/Basic.lean | 35 +++++++++++++++++++++++++++ 1 file changed, 35 insertions(+) diff --git a/Mathlib/FieldTheory/Finite/Basic.lean b/Mathlib/FieldTheory/Finite/Basic.lean index 59645c6b82ad7..134c9b94c12f6 100644 --- a/Mathlib/FieldTheory/Finite/Basic.lean +++ b/Mathlib/FieldTheory/Finite/Basic.lean @@ -526,6 +526,41 @@ theorem Int.ModEq.pow_card_sub_one_eq_one {p : ℕ} (hp : Nat.Prime p) {n : ℤ} · exact hpn.symm simpa [← ZMod.intCast_eq_intCast_iff] using ZMod.pow_card_sub_one_eq_one this +theorem pow_pow_modEq_one (p m a : ℕ) : (1 + p * a) ^ (p ^ m) ≡ 1 [MOD p ^ m] := by + induction' m with m hm + · exact Nat.modEq_one + · rw [Nat.ModEq.comm, add_comm, Nat.modEq_iff_dvd' (Nat.one_le_pow' _ _)] at hm + obtain ⟨d, hd⟩ := hm + rw [tsub_eq_iff_eq_add_of_le (Nat.one_le_pow' _ _), add_comm] at hd + rw [pow_succ, pow_mul, hd, add_pow, Finset.sum_range_succ', pow_zero, one_mul, one_pow, + one_mul, Nat.choose_zero_right, Nat.cast_one] + refine Nat.ModEq.add_right 1 (Nat.modEq_zero_iff_dvd.mpr ?_) + simp_rw [one_pow, mul_one, pow_succ', mul_assoc, ← Finset.mul_sum] + refine mul_dvd_mul_left (p ^ m) (dvd_mul_of_dvd_right (Finset.dvd_sum fun k hk ↦ ?_) d) + cases m + · rw [pow_zero, pow_one, one_mul, add_comm, add_left_inj] at hd + cases k <;> simp [← hd, mul_assoc, pow_succ'] + · cases k <;> simp [mul_assoc, pow_succ'] + +theorem ZMod.eq_one_or_isUnit_sub_one {n p k : ℕ} [Fact p.Prime] (hn : n = p ^ k) (a : ZMod n) + (ha : (orderOf a).Coprime n) : a = 1 ∨ IsUnit (a - 1) := by + rcases eq_or_ne n 0 with rfl | hn0 + · exact Or.inl (orderOf_eq_one_iff.mp ((orderOf a).coprime_zero_right.mp ha)) + rcases eq_or_ne a 0 with rfl | ha0 + · exact Or.inr (zero_sub (1 : ZMod n) ▸ isUnit_neg_one) + have : NeZero n := ⟨hn0⟩ + obtain ⟨a, rfl⟩ := ZMod.natCast_zmod_surjective a + rw [← orderOf_eq_one_iff, or_iff_not_imp_right] + refine fun h ↦ ha.eq_one_of_dvd ?_ + rw [orderOf_dvd_iff_pow_eq_one, ← Nat.cast_pow, ← Nat.cast_one, ZMod.eq_iff_modEq_nat, hn] + replace ha0 : 1 ≤ a := by + contrapose! ha0 + rw [Nat.lt_one_iff.mp ha0, Nat.cast_zero] + rw [← Nat.cast_one, ← Nat.cast_sub ha0, ZMod.isUnit_iff_coprime, hn] at h + obtain ⟨b, hb⟩ := not_imp_comm.mp (Nat.Prime.coprime_pow_of_not_dvd Fact.out) h + rw [tsub_eq_iff_eq_add_of_le ha0, add_comm] at hb + exact hb ▸ pow_pow_modEq_one p k b + section namespace FiniteField From 0662d0b6717c1777680219f9308baa68dd10abd8 Mon Sep 17 00:00:00 2001 From: Mitchell Lee Date: Sat, 14 Dec 2024 03:19:30 +0000 Subject: [PATCH 707/829] feat: more lemmas about the cocompact filter on topological spaces with a continuous multiplication and zero (#19650) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Let `M` be a topological space with a continuous multiplication operation and a zero. We prove the following lemmas. - `tendsto_mul_nhds_zero_prod_of_disjoint_cocompact`: Let `l` be a filter on `M` which is disjoint from the cocompact filter. Then the multiplication map `M × M → M` tends to zero on the filter product `𝓝 0 ×ˢ l`. (We also prove a variant with the order of the product swapped.) - `tendsto_mul_coprod_nhds_zero_inf_of_disjoint_cocompact`: Let `l` be a filter on `M × M` which is disjoint from the cocompact filter. Then the multiplication map `M × M → M` tends to zero on `Filter.coprod (𝓝 0) (𝓝 0) ⊓ l`. - `tendsto_mul_nhds_zero_of_disjoint_cocompact`: Let `l` be a filter on `M × M` which is disjoint from the cocompact filter and less than or equal to `Filter.coprod (𝓝 0) (𝓝 0)`. Then the multiplication map `M × M → M` tends to zero on `l`. - `Tendsto.tendsto_mul_zero_of_disjoint_cocompact`: Let `f : α → M` and `g : α → M` be functions. If `f` tends to zero on a filter `l` and the image of `l` under `g` is disjoint from the cocompact filter on `M`, then `fun (x : α) ↦ f x * g x` also tends to zero on `l`. (We also prove a variant with the order of the product swapped.) These lemmas were suggested by a reviewer comment on #12313: https://github.com/leanprover-community/mathlib4/pull/12313#issuecomment-2124740887. These lemmas are intended to simplify the proof of `tendsto_mul_cocompact_nhds_zero` in #12313, which states that if `f : α → M` and `g : β → M` are continuous functions that both tend to zero on the cocompact filter, then the "outer product" `fun (i : α × β) ↦ (f i.1) * (g i.2)` also tends to zero on the cocompact filter. --- Mathlib/Topology/Algebra/Monoid.lean | 77 +++++++++++++++++++++++++++- 1 file changed, 76 insertions(+), 1 deletion(-) diff --git a/Mathlib/Topology/Algebra/Monoid.lean b/Mathlib/Topology/Algebra/Monoid.lean index 2894a886145e8..e0485e870d3c6 100644 --- a/Mathlib/Topology/Algebra/Monoid.lean +++ b/Mathlib/Topology/Algebra/Monoid.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Johannes Hölzl, Mario Carneiro +Authors: Johannes Hölzl, Mario Carneiro, Mitchell Lee -/ import Mathlib.Algebra.BigOperators.Finprod import Mathlib.Order.Filter.Pointwise @@ -399,6 +399,81 @@ instance Submonoid.continuousMul [TopologicalSpace M] [Monoid M] [ContinuousMul (S : Submonoid M) : ContinuousMul S := S.toSubsemigroup.continuousMul +section MulZeroClass + +open Filter + +variable {α β : Type*} +variable [TopologicalSpace M] [MulZeroClass M] [ContinuousMul M] + +/-- Let `M` be a topological space with a continuous multiplication operation and a `0`. +Let `l` be a filter on `M` which is disjoint from the cocompact filter. Then, the multiplication map +`M × M → M` tends to zero on the filter product `𝓝 0 ×ˢ l`. -/ +theorem tendsto_mul_nhds_zero_prod_of_disjoint_cocompact {l : Filter M} + (hl : Disjoint l (cocompact M)) : + Tendsto (fun (x : M × M) ↦ x.1 * x.2) (𝓝 0 ×ˢ l) (𝓝 0) := calc + map (fun (x : M × M) ↦ x.1 * x.2) (𝓝 0 ×ˢ l) + _ ≤ map (fun (x : M × M) ↦ x.1 * x.2) (𝓝ˢ ({0} ×ˢ Set.univ)) := + map_mono <| nhds_prod_le_of_disjoint_cocompact 0 hl + _ ≤ 𝓝 0 := continuous_mul.tendsto_nhdsSet_nhds fun _ ⟨hx, _⟩ ↦ mul_eq_zero_of_left hx _ + +/-- Let `M` be a topological space with a continuous multiplication operation and a `0`. +Let `l` be a filter on `M` which is disjoint from the cocompact filter. Then, the multiplication map +`M × M → M` tends to zero on the filter product `l ×ˢ 𝓝 0`. -/ +theorem tendsto_mul_prod_nhds_zero_of_disjoint_cocompact {l : Filter M} + (hl : Disjoint l (cocompact M)) : + Tendsto (fun (x : M × M) ↦ x.1 * x.2) (l ×ˢ 𝓝 0) (𝓝 0) := calc + map (fun (x : M × M) ↦ x.1 * x.2) (l ×ˢ 𝓝 0) + _ ≤ map (fun (x : M × M) ↦ x.1 * x.2) (𝓝ˢ (Set.univ ×ˢ {0})) := + map_mono <| prod_nhds_le_of_disjoint_cocompact 0 hl + _ ≤ 𝓝 0 := continuous_mul.tendsto_nhdsSet_nhds fun _ ⟨_, hx⟩ ↦ mul_eq_zero_of_right _ hx + +/-- Let `M` be a topological space with a continuous multiplication operation and a `0`. +Let `l` be a filter on `M × M` which is disjoint from the cocompact filter. Then, the multiplication +map `M × M → M` tends to zero on `(𝓝 0).coprod (𝓝 0) ⊓ l`. -/ +theorem tendsto_mul_coprod_nhds_zero_inf_of_disjoint_cocompact {l : Filter (M × M)} + (hl : Disjoint l (cocompact (M × M))) : + Tendsto (fun (x : M × M) ↦ x.1 * x.2) ((𝓝 0).coprod (𝓝 0) ⊓ l) (𝓝 0) := by + have := calc + (𝓝 0).coprod (𝓝 0) ⊓ l + _ ≤ (𝓝 0).coprod (𝓝 0) ⊓ map Prod.fst l ×ˢ map Prod.snd l := + inf_le_inf_left _ le_prod_map_fst_snd + _ ≤ 𝓝 0 ×ˢ map Prod.snd l ⊔ map Prod.fst l ×ˢ 𝓝 0 := + Filter.coprod_inf_prod_le _ _ _ _ + apply (Tendsto.sup _ _).mono_left this + · apply tendsto_mul_nhds_zero_prod_of_disjoint_cocompact + exact disjoint_map_cocompact continuous_snd hl + · apply tendsto_mul_prod_nhds_zero_of_disjoint_cocompact + exact disjoint_map_cocompact continuous_fst hl + +/-- Let `M` be a topological space with a continuous multiplication operation and a `0`. +Let `l` be a filter on `M × M` which is both disjoint from the cocompact filter and less than or +equal to `(𝓝 0).coprod (𝓝 0)`. Then the multiplication map `M × M → M` tends to zero on `l`. -/ +theorem tendsto_mul_nhds_zero_of_disjoint_cocompact {l : Filter (M × M)} + (hl : Disjoint l (cocompact (M × M))) (h'l : l ≤ (𝓝 0).coprod (𝓝 0)) : + Tendsto (fun (x : M × M) ↦ x.1 * x.2) l (𝓝 0) := by + simpa [inf_eq_right.mpr h'l] using tendsto_mul_coprod_nhds_zero_inf_of_disjoint_cocompact hl + +/-- Let `M` be a topological space with a continuous multiplication operation and a `0`. +Let `f : α → M` and `g : α → M` be functions. If `f` tends to zero on a filter `l` +and the image of `l` under `g` is disjoint from the cocompact filter on `M`, then +`fun (x : α) ↦ f x * g x` also tends to zero on `l`. -/ +theorem Tendsto.tendsto_mul_zero_of_disjoint_cocompact_right {f g : α → M} {l : Filter α} + (hf : Tendsto f l (𝓝 0)) (hg : Disjoint (map g l) (cocompact M)) : + Tendsto (fun x ↦ f x * g x) l (𝓝 0) := + tendsto_mul_nhds_zero_prod_of_disjoint_cocompact hg |>.comp (hf.prod_mk tendsto_map) + +/-- Let `M` be a topological space with a continuous multiplication operation and a `0`. +Let `f : α → M` and `g : α → M` be functions. If `g` tends to zero on a filter `l` +and the image of `l` under `f` is disjoint from the cocompact filter on `M`, then +`fun (x : α) ↦ f x * g x` also tends to zero on `l`. -/ +theorem Tendsto.tendsto_mul_zero_of_disjoint_cocompact_left {f g : α → M} {l : Filter α} + (hf : Disjoint (map f l) (cocompact M)) (hg : Tendsto g l (𝓝 0)): + Tendsto (fun x ↦ f x * g x) l (𝓝 0) := + tendsto_mul_prod_nhds_zero_of_disjoint_cocompact hf |>.comp (tendsto_map.prod_mk hg) + +end MulZeroClass + section MulOneClass variable [TopologicalSpace M] [MulOneClass M] [ContinuousMul M] From ce0a54c9f5e4cc560d63dfca1b5ee12c024dd0f1 Mon Sep 17 00:00:00 2001 From: David Renshaw Date: Sat, 14 Dec 2024 10:44:26 +0000 Subject: [PATCH 708/829] perf: replace many instances of 'linarith' with 'omega' (#19951) Replaces 63 usages of `linarith` or `nlinarith` with `omega`. In all of these cases I have observed the new proof to be faster. I found these improvements by running [tryAtEachStep](https://github.com/dwrensha/tryAtEachStep) to try `omega` at every tactic step in mathlib. This is a continuation of the work from #11093. --- .../HomotopyCategory/DegreewiseSplit.lean | 2 +- .../Homology/HomotopyCategory/HomComplex.lean | 2 +- .../HomotopyCategory/HomComplexShift.lean | 4 ++-- Mathlib/Algebra/Lie/Weights/RootSystem.lean | 2 +- Mathlib/Algebra/Polynomial/Module/Basic.lean | 2 +- .../AlternatingFaceMapComplex.lean | 2 +- .../DoldKan/Degeneracies.lean | 3 +-- Mathlib/AlgebraicTopology/DoldKan/Faces.lean | 2 +- .../AlgebraicTopology/DoldKan/NCompGamma.lean | 2 +- .../Distribution/FourierSchwartz.lean | 2 +- .../Analysis/InnerProductSpace/TwoDim.lean | 2 +- Mathlib/Analysis/Normed/Algebra/Norm.lean | 2 +- .../Normed/Ring/SeminormFromConst.lean | 4 ++-- .../SpecialFunctions/Gamma/BohrMollerup.lean | 4 ++-- .../Triangulated/HomologicalFunctor.lean | 4 ++-- .../AkraBazzi/GrowsPolynomially.lean | 2 +- Mathlib/Data/Int/CardIntervalMod.lean | 4 ++-- Mathlib/Data/Real/Pi/Irrational.lean | 2 +- Mathlib/GroupTheory/ArchimedeanDensely.lean | 2 +- Mathlib/GroupTheory/Coxeter/Inversion.lean | 4 ++-- Mathlib/GroupTheory/Coxeter/Length.lean | 24 +++++++++---------- .../Eigenspace/Triangularizable.lean | 2 +- .../Measure/LevyProkhorovMetric.lean | 2 +- Mathlib/NumberTheory/AbelSummation.lean | 4 ++-- Mathlib/NumberTheory/Cyclotomic/Rat.lean | 2 +- Mathlib/NumberTheory/Fermat.lean | 2 +- Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean | 2 +- Mathlib/NumberTheory/Modular.lean | 4 ++-- .../NumberTheory/NumberField/Embeddings.lean | 6 ++--- Mathlib/NumberTheory/PythagoreanTriples.lean | 2 +- .../DedekindDomain/AdicValuation.lean | 2 +- Mathlib/RingTheory/DedekindDomain/Ideal.lean | 2 +- Mathlib/RingTheory/LaurentSeries.lean | 14 +++++------ Mathlib/RingTheory/PowerSeries/Basic.lean | 2 +- Mathlib/RingTheory/RingHom/Integral.lean | 2 +- 35 files changed, 62 insertions(+), 63 deletions(-) diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean b/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean index b9067c88df704..cabff6665b54e 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean @@ -115,7 +115,7 @@ noncomputable def mappingConeHomOfDegreewiseSplitIso : have s_g := (σ (p + 1)).s_g dsimp at r_f s_g ⊢ simp only [mappingConeHomOfDegreewiseSplitXIso, mappingCone.ext_from_iff _ _ _ rfl, - mappingCone.inl_v_d_assoc _ (p + 1) _ (p + 1 + 1) (by linarith) (by linarith), + mappingCone.inl_v_d_assoc _ (p + 1) _ (p + 1 + 1) (by linarith) (by omega), cocycleOfDegreewiseSplit, r_f, Int.reduceNeg, Cochain.ofHom_v, sub_comp, assoc, Hom.comm, comp_sub, mappingCone.inl_v_fst_v_assoc, mappingCone.inl_v_snd_v_assoc, shiftFunctor_obj_X', zero_comp, sub_zero, homOfDegreewiseSplit_f, diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/HomComplex.lean b/Mathlib/Algebra/Homology/HomotopyCategory/HomComplex.lean index d7de339fc2748..8707a528903f7 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/HomComplex.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/HomComplex.lean @@ -491,7 +491,7 @@ lemma δ_comp {n₁ n₂ n₁₂ : ℤ} (z₁ : Cochain F G n₁) (z₂ : Cochai dsimp rw [z₁.comp_v _ (add_assoc n₁ n₂ 1).symm p _ q rfl (by omega), Cochain.comp_v _ _ (show n₁ + 1 + n₂ = n₁ + n₂ + 1 by omega) p (p+n₁+1) q - (by linarith) (by omega), + (by omega) (by omega), δ_v (n₁ + n₂) _ rfl (z₁.comp z₂ rfl) p q hpq (p + n₁ + n₂) _ (by omega) rfl, z₁.comp_v z₂ rfl p _ _ rfl rfl, z₁.comp_v z₂ rfl (p+1) (p+n₁+1) q (by omega) (by omega), diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/HomComplexShift.lean b/Mathlib/Algebra/Homology/HomotopyCategory/HomComplexShift.lean index be8303566c6d1..647dcd6a0dffb 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/HomComplexShift.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/HomComplexShift.lean @@ -85,7 +85,7 @@ lemma rightUnshift_v {n' a : ℤ} (γ : Cochain K (L⟦a⟧) n') (n : ℤ) (hn : def leftUnshift {n' a : ℤ} (γ : Cochain (K⟦a⟧) L n') (n : ℤ) (hn : n + a = n') : Cochain K L n := Cochain.mk (fun p q hpq => (a * n' + ((a * (a-1))/2)).negOnePow • - (K.shiftFunctorObjXIso a (p - a) p (by linarith)).inv ≫ γ.v (p-a) q (by omega)) + (K.shiftFunctorObjXIso a (p - a) p (by omega)).inv ≫ γ.v (p-a) q (by omega)) lemma leftUnshift_v {n' a : ℤ} (γ : Cochain (K⟦a⟧) L n') (n : ℤ) (hn : n + a = n') (p q : ℤ) (hpq : p + n = q) (p' : ℤ) (hp' : p' + n' = q) : @@ -373,7 +373,7 @@ lemma leftShift_comp (a n' : ℤ) (hn' : n + a = n') {m t t' : ℤ} (γ' : Cocha (γ.comp γ' h).leftShift a t' ht' = (a * m).negOnePow • (γ.leftShift a n' hn').comp γ' (by rw [← ht', ← h, ← hn', add_assoc, add_comm a, add_assoc]) := by ext p q hpq - have h' : n' + m = t' := by linarith + have h' : n' + m = t' := by omega dsimp simp only [Cochain.comp_v _ _ h' p (p + n') q rfl (by omega), γ.leftShift_v a n' hn' p (p + n') rfl (p + a) (by omega), diff --git a/Mathlib/Algebra/Lie/Weights/RootSystem.lean b/Mathlib/Algebra/Lie/Weights/RootSystem.lean index c6334632c472d..052e935cc4836 100644 --- a/Mathlib/Algebra/Lie/Weights/RootSystem.lean +++ b/Mathlib/Algebra/Lie/Weights/RootSystem.lean @@ -119,7 +119,7 @@ lemma rootSpace_neg_nsmul_add_chainTop_of_lt (hα : α.IsNonZero) {n : ℕ} (hn neg_add_rev, apply_coroot_eq_cast' α β, Int.cast_sub, Int.cast_mul, Int.cast_ofNat, mul_comm (2 : K), add_sub_cancel, neg_neg, add_sub, Nat.cast_inj, eq_sub_iff_add_eq, ← Nat.cast_add, ← sub_eq_neg_add, sub_eq_iff_eq_add] at this - linarith [this, hn] + omega have H₂ : ((1 + n + chainTopCoeff (-α) W) • α + chainTop (-α) W : H → K) = (chainTopCoeff α β + 1) • α + β := by simp only [Weight.coe_neg, ← Nat.cast_smul_eq_nsmul ℤ, Nat.cast_add, Nat.cast_one, coe_chainTop, diff --git a/Mathlib/Algebra/Polynomial/Module/Basic.lean b/Mathlib/Algebra/Polynomial/Module/Basic.lean index 047d5c7f86864..ce68ea47f409d 100644 --- a/Mathlib/Algebra/Polynomial/Module/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Module/Basic.lean @@ -153,7 +153,7 @@ theorem smul_single_apply (i : ℕ) (f : R[X]) (m : M) (n : ℕ) : · rw [if_neg h, ite_eq_right_iff] intro e exfalso - linarith + omega theorem smul_apply (f : R[X]) (g : PolynomialModule R M) (n : ℕ) : (f • g) n = ∑ x ∈ Finset.antidiagonal n, f.coeff x.1 • g x.2 := by diff --git a/Mathlib/AlgebraicTopology/AlternatingFaceMapComplex.lean b/Mathlib/AlgebraicTopology/AlternatingFaceMapComplex.lean index 7e25833d9f89f..364ebb8abf7fb 100644 --- a/Mathlib/AlgebraicTopology/AlternatingFaceMapComplex.lean +++ b/Mathlib/AlgebraicTopology/AlternatingFaceMapComplex.lean @@ -82,7 +82,7 @@ theorem d_squared (n : ℕ) : objD X (n + 1) ≫ objD X n = 0 := by intro ij hij simp only [S, φ, Finset.mem_univ, Finset.compl_filter, Finset.mem_filter, true_and, Fin.val_succ, Fin.coe_castLT] at hij ⊢ - linarith + omega · -- φ : S → Sᶜ is injective rintro ⟨i, j⟩ hij ⟨i', j'⟩ hij' h rw [Prod.mk.inj_iff] diff --git a/Mathlib/AlgebraicTopology/DoldKan/Degeneracies.lean b/Mathlib/AlgebraicTopology/DoldKan/Degeneracies.lean index dde1b3dea55c9..3b60b11d9749d 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/Degeneracies.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/Degeneracies.lean @@ -56,8 +56,7 @@ theorem σ_comp_P_eq_zero (X : SimplicialObject C) {n q : ℕ} (i : Fin (n + 1)) revert i hi induction' q with q hq · intro i (hi : n + 1 ≤ i) - exfalso - linarith [Fin.is_lt i] + omega · intro i (hi : n + 1 ≤ i + q + 1) by_cases h : n + 1 ≤ (i : ℕ) + q · rw [P_succ, HomologicalComplex.comp_f, ← assoc, hq i h, zero_comp] diff --git a/Mathlib/AlgebraicTopology/DoldKan/Faces.lean b/Mathlib/AlgebraicTopology/DoldKan/Faces.lean index 845dd32f779d1..3a46ea7d170ba 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/Faces.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/Faces.lean @@ -196,7 +196,7 @@ theorem induction {Y : C} {n q : ℕ} {φ : Y ⟶ X _[n + 1]} (v : HigherFacesVa swap · rw [Fin.le_iff_val_le_val] dsimp - linarith + omega simp only [← assoc, v j (by omega), zero_comp] · -- in the last case, a=m, q=1 and j=a+1 rw [X.δ_comp_δ_self'_assoc] diff --git a/Mathlib/AlgebraicTopology/DoldKan/NCompGamma.lean b/Mathlib/AlgebraicTopology/DoldKan/NCompGamma.lean index a62ceb55bcb32..05cf1193af780 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/NCompGamma.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/NCompGamma.lean @@ -73,7 +73,7 @@ theorem PInfty_comp_map_mono_eq_zero (X : SimplicialObject C) {n : ℕ} {Δ' : S · simp only [op_comp, X.map_comp, assoc, PInfty_f] erw [(HigherFacesVanish.of_P _ _).comp_δ_eq_zero_assoc _ hj₁, zero_comp] by_contra - exact hj₁ (by simp only [Fin.ext_iff, Fin.val_zero]; linarith) + exact hj₁ (by simp only [Fin.ext_iff, Fin.val_zero]; omega) @[reassoc] theorem Γ₀_obj_termwise_mapMono_comp_PInfty (X : SimplicialObject C) {Δ Δ' : SimplexCategory} diff --git a/Mathlib/Analysis/Distribution/FourierSchwartz.lean b/Mathlib/Analysis/Distribution/FourierSchwartz.lean index e3a6bac433ffe..3c7fb985a3f49 100644 --- a/Mathlib/Analysis/Distribution/FourierSchwartz.lean +++ b/Mathlib/Analysis/Distribution/FourierSchwartz.lean @@ -71,7 +71,7 @@ noncomputable def fourierTransformCLM : 𝓢(V, E) →L[𝕜] 𝓢(V, E) := by have : (p.1 + integrablePower (volume : Measure V), p.2) ∈ (Finset.range (n + integrablePower (volume : Measure V) + 1) ×ˢ Finset.range (k + 1)) := by simp [hp.2] - linarith + omega apply Finset.le_sup this (f := fun p ↦ SchwartzMap.seminorm 𝕜 p.1 p.2 (E := V) (F := E)) _ = _ := by simp [mul_assoc] diff --git a/Mathlib/Analysis/InnerProductSpace/TwoDim.lean b/Mathlib/Analysis/InnerProductSpace/TwoDim.lean index 18f8a99312e5f..a9afb01f2a21b 100644 --- a/Mathlib/Analysis/InnerProductSpace/TwoDim.lean +++ b/Mathlib/Analysis/InnerProductSpace/TwoDim.lean @@ -211,7 +211,7 @@ def rightAngleRotationAux₂ : E →ₗᵢ[ℝ] E := have : Finset.card {x} = 1 := Finset.card_singleton x have : finrank ℝ K + finrank ℝ Kᗮ = finrank ℝ E := K.finrank_add_finrank_orthogonal have : finrank ℝ E = 2 := Fact.out - linarith + omega obtain ⟨w, hw₀⟩ : ∃ w : Kᗮ, w ≠ 0 := exists_ne 0 have hw' : ⟪x, (w : E)⟫ = 0 := Submodule.mem_orthogonal_singleton_iff_inner_right.mp w.2 have hw : (w : E) ≠ 0 := fun h => hw₀ (Submodule.coe_eq_zero.mp h) diff --git a/Mathlib/Analysis/Normed/Algebra/Norm.lean b/Mathlib/Analysis/Normed/Algebra/Norm.lean index 715ad8694cf6f..bcfc5b377e154 100644 --- a/Mathlib/Analysis/Normed/Algebra/Norm.lean +++ b/Mathlib/Analysis/Normed/Algebra/Norm.lean @@ -188,7 +188,7 @@ def toRingNorm (f : MulRingNorm R) : RingNorm R where /-- 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 + · omega · rw [map_pow] end MulRingNorm diff --git a/Mathlib/Analysis/Normed/Ring/SeminormFromConst.lean b/Mathlib/Analysis/Normed/Ring/SeminormFromConst.lean index 91ff0ab51079c..46c9dbc76ed58 100644 --- a/Mathlib/Analysis/Normed/Ring/SeminormFromConst.lean +++ b/Mathlib/Analysis/Normed/Ring/SeminormFromConst.lean @@ -145,7 +145,7 @@ def seminormFromConst : RingSeminorm R where apply (seminormFromConst_isLimit hf1 hc hpm (x * y)).comp (tendsto_atTop_atTop_of_monotone (fun _ _ hnm ↦ by simp only [mul_le_mul_left, Nat.succ_pos', hnm]) _) - · rintro n; use n; linarith + · rintro n; use n; omega refine le_of_tendsto_of_tendsto' hlim ((seminormFromConst_isLimit hf1 hc hpm x).mul (seminormFromConst_isLimit hf1 hc hpm y)) (fun n ↦ ?_) simp only [seminormFromConst_seq] @@ -246,7 +246,7 @@ theorem seminormFromConst_const_mul (x : R) : (𝓝 (seminormFromConst' hf1 hc hpm x)) := by apply (seminormFromConst_isLimit hf1 hc hpm x).comp (tendsto_atTop_atTop_of_monotone (fun _ _ hnm ↦ add_le_add_right hnm 1) _) - rintro n; use n; linarith + rintro n; use n; omega rw [seminormFromConst_apply_c hf1 hc hpm] apply tendsto_nhds_unique (seminormFromConst_isLimit hf1 hc hpm (c * x)) have hterm : seminormFromConst_seq c f (c * x) = diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/BohrMollerup.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/BohrMollerup.lean index 6a9a40ef2d4ff..ef7ef15bafc12 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gamma/BohrMollerup.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/BohrMollerup.lean @@ -176,7 +176,7 @@ theorem f_add_nat_le (hf_conv : ConvexOn ℝ (Ioi 0) f) theorem f_add_nat_ge (hf_conv : ConvexOn ℝ (Ioi 0) f) (hf_feq : ∀ {y : ℝ}, 0 < y → f (y + 1) = f y + log y) (hn : 2 ≤ n) (hx : 0 < x) : f n + x * log (n - 1) ≤ f (n + x) := by - have npos : 0 < (n : ℝ) - 1 := by rw [← Nat.cast_one, sub_pos, Nat.cast_lt]; linarith + have npos : 0 < (n : ℝ) - 1 := by rw [← Nat.cast_one, sub_pos, Nat.cast_lt]; omega have c := (convexOn_iff_slope_mono_adjacent.mp <| hf_conv).2 npos (by linarith : 0 < (n : ℝ) + x) (by linarith : (n : ℝ) - 1 < (n : ℝ)) (by linarith) @@ -209,7 +209,7 @@ theorem le_logGammaSeq (hf_conv : ConvexOn ℝ (Ioi 0) f) f x ≤ f 1 + x * log (n + 1) - x * log n + logGammaSeq x n := by rw [logGammaSeq, ← add_sub_assoc, le_sub_iff_add_le, ← f_add_nat_eq (@hf_feq) hx, add_comm x] refine (f_add_nat_le hf_conv (@hf_feq) (Nat.add_one_ne_zero n) hx hx').trans (le_of_eq ?_) - rw [f_nat_eq @hf_feq (by linarith : n + 1 ≠ 0), Nat.add_sub_cancel, Nat.cast_add_one] + rw [f_nat_eq @hf_feq (by omega : n + 1 ≠ 0), Nat.add_sub_cancel, Nat.cast_add_one] ring theorem ge_logGammaSeq (hf_conv : ConvexOn ℝ (Ioi 0) f) diff --git a/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean b/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean index 5c1911f19648b..afa7602914d26 100644 --- a/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean +++ b/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean @@ -215,12 +215,12 @@ lemma homologySequence_exact₃ : (ShortComplex.mk _ _ (F.comp_homologySequenceδ T hT _ _ h)).Exact := by refine ShortComplex.exact_of_iso ?_ (F.homologySequence_exact₂ _ (rot_of_distTriang _ hT) n₀) exact ShortComplex.isoMk (Iso.refl _) (Iso.refl _) - ((F.shiftIso 1 n₀ n₁ (by linarith)).app _) (by simp) (by simp [homologySequenceδ, shiftMap]) + ((F.shiftIso 1 n₀ n₁ (by omega)).app _) (by simp) (by simp [homologySequenceδ, shiftMap]) lemma homologySequence_exact₁ : (ShortComplex.mk _ _ (F.homologySequenceδ_comp T hT _ _ h)).Exact := by refine ShortComplex.exact_of_iso ?_ (F.homologySequence_exact₂ _ (inv_rot_of_distTriang _ hT) n₁) - refine ShortComplex.isoMk (-((F.shiftIso (-1) n₁ n₀ (by linarith)).app _)) + refine ShortComplex.isoMk (-((F.shiftIso (-1) n₁ n₀ (by omega)).app _)) (Iso.refl _) (Iso.refl _) ?_ (by simp) dsimp simp only [homologySequenceδ, neg_comp, map_neg, comp_id, diff --git a/Mathlib/Computability/AkraBazzi/GrowsPolynomially.lean b/Mathlib/Computability/AkraBazzi/GrowsPolynomially.lean index 98cc54a85dcff..9a1922aa8af6a 100644 --- a/Mathlib/Computability/AkraBazzi/GrowsPolynomially.lean +++ b/Mathlib/Computability/AkraBazzi/GrowsPolynomially.lean @@ -114,7 +114,7 @@ lemma eventually_zero_of_frequently_zero (hf : GrowsPolynomially f) (hf' : ∃ rw [Set.left_mem_Icc] gcongr · norm_num - · linarith + · omega simp only [ih, mul_zero, Set.Icc_self, Set.mem_singleton_iff] at hx refine hx ⟨?lb₁, ?ub₁⟩ case lb₁ => diff --git a/Mathlib/Data/Int/CardIntervalMod.lean b/Mathlib/Data/Int/CardIntervalMod.lean index b243bdf2208d0..3fc6f73fcef99 100644 --- a/Mathlib/Data/Int/CardIntervalMod.lean +++ b/Mathlib/Data/Int/CardIntervalMod.lean @@ -91,7 +91,7 @@ lemma Ico_filter_modEq_cast {v : ℕ} : simp only [mem_map, mem_filter, mem_Ico, castEmbedding_apply] constructor · simp_rw [forall_exists_index, ← natCast_modEq_iff]; intro y ⟨h, c⟩; subst c; exact_mod_cast h - · intro h; lift x to ℕ using (by linarith); exact ⟨x, by simp_all [natCast_modEq_iff]⟩ + · intro h; lift x to ℕ using (by omega); exact ⟨x, by simp_all [natCast_modEq_iff]⟩ lemma Ioc_filter_modEq_cast {v : ℕ} : {x ∈ Ioc a b | x ≡ v [MOD r]}.map castEmbedding = @@ -100,7 +100,7 @@ lemma Ioc_filter_modEq_cast {v : ℕ} : simp only [mem_map, mem_filter, mem_Ioc, castEmbedding_apply] constructor · simp_rw [forall_exists_index, ← natCast_modEq_iff]; intro y ⟨h, c⟩; subst c; exact_mod_cast h - · intro h; lift x to ℕ using (by linarith); exact ⟨x, by simp_all [natCast_modEq_iff]⟩ + · intro h; lift x to ℕ using (by omega); exact ⟨x, by simp_all [natCast_modEq_iff]⟩ variable (hr : 0 < r) include hr diff --git a/Mathlib/Data/Real/Pi/Irrational.lean b/Mathlib/Data/Real/Pi/Irrational.lean index 34d2652435e86..fbecb8f749514 100644 --- a/Mathlib/Data/Real/Pi/Irrational.lean +++ b/Mathlib/Data/Real/Pi/Irrational.lean @@ -291,7 +291,7 @@ private lemma not_irrational_exists_rep {x : ℝ} : 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)) + is_integer a b ((sinPoly_natDegree_le _).trans (by omega)) 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) = diff --git a/Mathlib/GroupTheory/ArchimedeanDensely.lean b/Mathlib/GroupTheory/ArchimedeanDensely.lean index 6e65e049cb576..27eb25a932009 100644 --- a/Mathlib/GroupTheory/ArchimedeanDensely.lean +++ b/Mathlib/GroupTheory/ArchimedeanDensely.lean @@ -207,7 +207,7 @@ lemma LinearOrderedAddCommGroup.discrete_iff_not_denselyOrdered (G : Type*) intro e H rw [denselyOrdered_iff_of_orderIsoClass e] at H obtain ⟨_, _⟩ := exists_between (one_pos (α := ℤ)) - linarith + omega variable (G) in /-- Any linearly ordered mul-archimedean group is either isomorphic (and order-isomorphic) diff --git a/Mathlib/GroupTheory/Coxeter/Inversion.lean b/Mathlib/GroupTheory/Coxeter/Inversion.lean index b757715af2158..43a790543f337 100644 --- a/Mathlib/GroupTheory/Coxeter/Inversion.lean +++ b/Mathlib/GroupTheory/Coxeter/Inversion.lean @@ -455,7 +455,7 @@ theorem IsReduced.nodup_leftInvSeq {ω : List B} (rω : cs.IsReduced ω) : List. lemma getElem_succ_leftInvSeq_alternatingWord (i j : B) (p k : ℕ) (h : k + 1 < 2 * p) : (lis (alternatingWord i j (2 * p)))[k + 1]'(by simp; exact h) = - MulAut.conj (s i) ((lis (alternatingWord j i (2 * p)))[k]'(by simp; linarith)) := by + MulAut.conj (s i) ((lis (alternatingWord j i (2 * p)))[k]'(by simp; omega)) := by rw [cs.getElem_leftInvSeq (alternatingWord i j (2 * p)) (k + 1) (by simp[h]), cs.getElem_leftInvSeq (alternatingWord j i (2 * p)) k (by simp[h]; omega)] simp only [MulAut.conj, listTake_succ_alternatingWord i j p k h, cs.wordProd_cons, mul_assoc, @@ -466,7 +466,7 @@ lemma getElem_succ_leftInvSeq_alternatingWord theorem getElem_leftInvSeq_alternatingWord (i j : B) (p k : ℕ) (h : k < 2 * p) : - (lis (alternatingWord i j (2 * p)))[k]'(by simp; linarith) = + (lis (alternatingWord i j (2 * p)))[k]'(by simp; omega) = π alternatingWord j i (2 * k + 1) := by revert i j induction k with diff --git a/Mathlib/GroupTheory/Coxeter/Length.lean b/Mathlib/GroupTheory/Coxeter/Length.lean index da1191305642b..3cd6c22bf33c7 100644 --- a/Mathlib/GroupTheory/Coxeter/Length.lean +++ b/Mathlib/GroupTheory/Coxeter/Length.lean @@ -183,13 +183,13 @@ theorem length_mul_simple (w : W) (i : B) : have length_ge := cs.length_mul_ge_length_sub_length w (s i) simp only [length_simple, tsub_le_iff_right] at length_ge -- length_ge : ℓ w ≤ ℓ (w * s i) + 1 - linarith + omega · -- gt : ℓ w < ℓ (w * s i) left have length_le := cs.length_mul_le w (s i) simp only [length_simple] at length_le -- length_le : ℓ (w * s i) ≤ ℓ w + 1 - linarith + omega theorem length_simple_mul (w : W) (i : B) : ℓ (s i * w) = ℓ w + 1 ∨ ℓ (s i * w) + 1 = ℓ w := by @@ -222,7 +222,7 @@ private theorem isReduced_take_and_drop {ω : List B} (hω : cs.IsReduced ω) (j _ = ℓ (π (ω.take j) * π (ω.drop j)) := by rw [← cs.wordProd_append, ω.take_append_drop j] _ ≤ ℓ (π (ω.take j)) + ℓ (π (ω.drop j)) := cs.length_mul_le _ _ unfold IsReduced - exact ⟨by linarith, by linarith⟩ + omega theorem isReduced_take {ω : List B} (hω : cs.IsReduced ω) (j : ℕ) : cs.IsReduced (ω.take j) := (isReduced_take_and_drop _ hω _).1 @@ -237,7 +237,7 @@ theorem not_isReduced_alternatingWord (i i' : B) {m : ℕ} (hM : M i i' ≠ 0) ( suffices h : ℓ (π (alternatingWord i i' (M i i' + 1))) < M i i' + 1 by unfold IsReduced rw [Nat.succ_eq_add_one, length_alternatingWord] - linarith + omega have : M i i' + 1 ≤ M i i' * 2 := by linarith [Nat.one_le_iff_ne_zero.mpr hM] rw [cs.prod_alternatingWord_eq_prod_alternatingWord_sub i i' _ this] have : M i i' * 2 - (M i i' + 1) = M i i' - 1 := by @@ -299,36 +299,36 @@ theorem isLeftDescent_iff {w : W} {i : B} : unfold IsLeftDescent constructor · intro _ - exact (cs.length_simple_mul w i).resolve_left (by linarith) + exact (cs.length_simple_mul w i).resolve_left (by omega) · intro _ - linarith + omega theorem not_isLeftDescent_iff {w : W} {i : B} : ¬cs.IsLeftDescent w i ↔ ℓ (s i * w) = ℓ w + 1 := by unfold IsLeftDescent constructor · intro _ - exact (cs.length_simple_mul w i).resolve_right (by linarith) + exact (cs.length_simple_mul w i).resolve_right (by omega) · intro _ - linarith + omega theorem isRightDescent_iff {w : W} {i : B} : cs.IsRightDescent w i ↔ ℓ (w * s i) + 1 = ℓ w := by unfold IsRightDescent constructor · intro _ - exact (cs.length_mul_simple w i).resolve_left (by linarith) + exact (cs.length_mul_simple w i).resolve_left (by omega) · intro _ - linarith + omega theorem not_isRightDescent_iff {w : W} {i : B} : ¬cs.IsRightDescent w i ↔ ℓ (w * s i) = ℓ w + 1 := by unfold IsRightDescent constructor · intro _ - exact (cs.length_mul_simple w i).resolve_right (by linarith) + exact (cs.length_mul_simple w i).resolve_right (by omega) · intro _ - linarith + omega theorem isLeftDescent_iff_not_isLeftDescent_mul {w : W} {i : B} : cs.IsLeftDescent w i ↔ ¬cs.IsLeftDescent (s i * w) i := by diff --git a/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean b/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean index 6a90c71055221..fcb07a9d441a2 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean @@ -107,7 +107,7 @@ theorem iSup_maxGenEigenspace_eq_top [IsAlgClosed K] [FiniteDimensional K V] (f rw [Module.End.genEigenspace_nat, Module.End.genEigenrange_nat] apply LinearMap.finrank_range_add_finrank_ker -- Therefore the dimension `ER` mus be smaller than `finrank K V`. - have h_dim_ER : finrank K ER < n.succ := by linarith + have h_dim_ER : finrank K ER < n.succ := by omega -- This allows us to apply the induction hypothesis on `ER`: have ih_ER : ⨆ (μ : K), f'.maxGenEigenspace μ = ⊤ := ih (finrank K ER) h_dim_ER f' rfl diff --git a/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean b/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean index 3fb7d9b953875..ecb9985af6a9f 100644 --- a/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean +++ b/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean @@ -614,7 +614,7 @@ lemma LevyProkhorov.continuous_equiv_symm_probabilityMeasure : simp only [mem_Iio, compl_iUnion, mem_iInter, mem_compl_iff, not_forall, not_not, exists_prop] at con obtain ⟨j, j_small, ω_in_Esj⟩ := con - exact disjoint_left.mp (Es_disjoint (show j ≠ i by linarith)) ω_in_Esj ω_in_Esi + exact disjoint_left.mp (Es_disjoint (show j ≠ i by omega)) ω_in_Esj ω_in_Esi intro ω ω_in_B obtain ⟨i, hi⟩ := show ∃ n, ω ∈ Es n by simp only [← mem_iUnion, Es_cover, mem_univ] simp only [mem_Ici, mem_union, mem_iUnion, exists_prop] diff --git a/Mathlib/NumberTheory/AbelSummation.lean b/Mathlib/NumberTheory/AbelSummation.lean index 69dc93ce98099..843b01ff64820 100644 --- a/Mathlib/NumberTheory/AbelSummation.lean +++ b/Mathlib/NumberTheory/AbelSummation.lean @@ -162,8 +162,8 @@ theorem sum_mul_eq_sub_integral_mul' (hc : c 0 = 0) (b : ℝ) f b * (∑ k ∈ Icc 0 ⌊b⌋₊, c k) - ∫ t in Set.Ioc 1 b, deriv f t * (∑ k ∈ Icc 0 ⌊t⌋₊, c k) := by obtain hb | hb := le_or_gt 1 b · have : 1 ≤ ⌊b⌋₊ := (Nat.one_le_floor_iff _).mpr hb - nth_rewrite 1 [Finset.Icc_eq_cons_Ioc (by linarith), sum_cons, ← Nat.Icc_succ_left, - Finset.Icc_eq_cons_Ioc (by linarith), sum_cons] + nth_rewrite 1 [Finset.Icc_eq_cons_Ioc (by omega), sum_cons, ← Nat.Icc_succ_left, + Finset.Icc_eq_cons_Ioc (by omega), sum_cons] rw [Nat.succ_eq_add_one, zero_add, ← Nat.floor_one (α := ℝ), sum_mul_eq_sub_sub_integral_mul c zero_le_one hb hf_diff hf_int, Nat.floor_one, Nat.cast_one, Finset.Icc_eq_cons_Ioc zero_le_one, sum_cons, show 1 = 0 + 1 by rfl, Nat.Ioc_succ_singleton, diff --git a/Mathlib/NumberTheory/Cyclotomic/Rat.lean b/Mathlib/NumberTheory/Cyclotomic/Rat.lean index 2c898d2ea3c30..dad9eca2e6965 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Rat.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Rat.lean @@ -474,7 +474,7 @@ theorem not_exists_int_prime_dvd_sub_of_prime_pow_ne_two simp only [↓reduceIte, map_add, Finsupp.coe_add, Pi.add_apply] at h rw [show (p : 𝓞 K) * x = (p : ℤ) • x by simp, ← pB.basis.coord_apply, LinearMap.map_smul, ← zsmul_one, ← pB.basis.coord_apply, LinearMap.map_smul, - show 1 = pB.gen ^ (⟨0, by linarith⟩ : Fin pB.dim).1 by simp, ← pB.basis_eq_pow, + show 1 = pB.gen ^ (⟨0, by omega⟩ : Fin pB.dim).1 by simp, ← pB.basis_eq_pow, pB.basis.coord_apply, pB.basis.coord_apply, pB.basis.repr_self_apply] at h simp only [smul_eq_mul, Fin.mk.injEq, zero_ne_one, ↓reduceIte, mul_zero, add_zero] at h exact (Int.prime_iff_natAbs_prime.2 (by simp [hp.1])).not_dvd_one ⟨_, h⟩ diff --git a/Mathlib/NumberTheory/Fermat.lean b/Mathlib/NumberTheory/Fermat.lean index c2aca1d62719d..cd922700db163 100644 --- a/Mathlib/NumberTheory/Fermat.lean +++ b/Mathlib/NumberTheory/Fermat.lean @@ -76,7 +76,7 @@ theorem two_mul_fermatNumber_sub_one_sq_le_fermatNumber_sq (n : ℕ) : simp only [fermatNumber, add_tsub_cancel_right] have : 0 ≤ 1 + 2 ^ (2 ^ n * 4) := le_add_left 0 (Nat.add 1 _) ring_nf - linarith + omega theorem fermatNumber_eq_fermatNumber_sq_sub_two_mul_fermatNumber_sub_one_sq (n : ℕ) : fermatNumber (n + 2) = (fermatNumber (n + 1)) ^ 2 - 2 * (fermatNumber n - 1) ^ 2 := by diff --git a/Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean b/Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean index ae321a0cce2eb..92a7baa1debe2 100644 --- a/Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean +++ b/Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean @@ -258,7 +258,7 @@ lemma continuousOn_term (n : ℕ) : · positivity · exact rpow_le_rpow_of_exponent_le (le_trans (by simp) hx.1.le) (by linarith) · rw [← IntegrableOn, ← intervalIntegrable_iff_integrableOn_Ioc_of_le (by linarith)] - exact_mod_cast term_welldef (by linarith : 0 < (n + 1)) zero_lt_one + exact_mod_cast term_welldef (by omega : 0 < (n + 1)) zero_lt_one · rw [ae_restrict_iff' measurableSet_Ioc] filter_upwards with x hx refine continuousOn_of_forall_continuousAt (fun s (hs : 1 ≤ s) ↦ continuousAt_const.div ?_ ?_) diff --git a/Mathlib/NumberTheory/Modular.lean b/Mathlib/NumberTheory/Modular.lean index acad10c21c22c..8b2189779e182 100644 --- a/Mathlib/NumberTheory/Modular.lean +++ b/Mathlib/NumberTheory/Modular.lean @@ -324,7 +324,7 @@ variable {z} theorem exists_eq_T_zpow_of_c_eq_zero (hc : g 1 0 = 0) : ∃ n : ℤ, ∀ z : ℍ, g • z = T ^ n • z := by have had := g.det_coe - replace had : g 0 0 * g 1 1 = 1 := by rw [det_fin_two, hc] at had; linarith + replace had : g 0 0 * g 1 1 = 1 := by rw [det_fin_two, hc] at had; omega rcases Int.eq_one_or_neg_one_of_mul_eq_one' had with (⟨ha, hd⟩ | ⟨ha, hd⟩) · use g 0 1 suffices g = T ^ g 0 1 by intro z; conv_lhs => rw [this] @@ -338,7 +338,7 @@ theorem exists_eq_T_zpow_of_c_eq_zero (hc : g 1 0 = 0) : -- If `c = 1`, then `g` factorises into a product terms involving only `T` and `S`. theorem g_eq_of_c_eq_one (hc : g 1 0 = 1) : g = T ^ g 0 0 * S * T ^ g 1 1 := by have hg := g.det_coe.symm - replace hg : g 0 1 = g 0 0 * g 1 1 - 1 := by rw [det_fin_two, hc] at hg; linarith + replace hg : g 0 1 = g 0 0 * g 1 1 - 1 := by rw [det_fin_two, hc] at hg; omega refine Subtype.ext ?_ conv_lhs => rw [(g : Matrix _ _ ℤ).eta_fin_two] simp only [hg, sub_eq_add_neg, hc, coe_mul, coe_T_zpow, coe_S, mul_fin_two, mul_zero, mul_one, diff --git a/Mathlib/NumberTheory/NumberField/Embeddings.lean b/Mathlib/NumberTheory/NumberField/Embeddings.lean index 9cda9eb7b84d8..1d5de4b41ece5 100644 --- a/Mathlib/NumberTheory/NumberField/Embeddings.lean +++ b/Mathlib/NumberTheory/NumberField/Embeddings.lean @@ -1061,16 +1061,16 @@ theorem nrRealPlaces_eq_zero_of_two_lt (hk : 2 < k) (hζ : IsPrimitiveRoot ζ k) congr have hre : (f ζ).re = 1 ∨ (f ζ).re = -1 := by rw [← Complex.abs_re_eq_abs] at him - have := Complex.norm_eq_one_of_pow_eq_one hζ'.pow_eq_one (by linarith) + have := Complex.norm_eq_one_of_pow_eq_one hζ'.pow_eq_one (by omega) rwa [Complex.norm_eq_abs, ← him, ← abs_one, abs_eq_abs] at this cases hre with | inl hone => - exact hζ'.ne_one (by linarith) <| Complex.ext (by simp [hone]) (by simp [him]) + exact hζ'.ne_one (by omega) <| Complex.ext (by simp [hone]) (by simp [him]) | inr hnegone => replace hζ' := hζ'.eq_orderOf simp only [show f ζ = -1 from Complex.ext (by simp [hnegone]) (by simp [him]), orderOf_neg_one, ringChar.eq_zero, OfNat.zero_ne_ofNat, ↓reduceIte] at hζ' - linarith + omega end IsPrimitiveRoot diff --git a/Mathlib/NumberTheory/PythagoreanTriples.lean b/Mathlib/NumberTheory/PythagoreanTriples.lean index 1b118d57fb50c..d4c366db72320 100644 --- a/Mathlib/NumberTheory/PythagoreanTriples.lean +++ b/Mathlib/NumberTheory/PythagoreanTriples.lean @@ -634,7 +634,7 @@ theorem coprime_classification' {x y z : ℤ} (h : PythagoreanTriple x y z) exact ht3 · rw [Int.neg_emod_two, Int.neg_emod_two] apply And.intro ht4 - linarith + omega · exfalso revert h_pos rw [h_neg] diff --git a/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean b/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean index 9afbb4385a989..0aac9bcf5b559 100644 --- a/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean +++ b/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean @@ -538,6 +538,6 @@ lemma adicCompletion.mul_nonZeroDivisor_mem_adicCompletionIntegers (v : HeightOn Int.natCast_natAbs, smul_eq_mul], ← Int.eq_natAbs_of_zero_le ha.le, smul_eq_mul] -- and now it's easy - linarith + omega end IsDedekindDomain.HeightOneSpectrum diff --git a/Mathlib/RingTheory/DedekindDomain/Ideal.lean b/Mathlib/RingTheory/DedekindDomain/Ideal.lean index 8ad308eb008d5..acc74b557241b 100644 --- a/Mathlib/RingTheory/DedekindDomain/Ideal.lean +++ b/Mathlib/RingTheory/DedekindDomain/Ideal.lean @@ -1192,7 +1192,7 @@ theorem count_associates_factors_eq [DecidableEq (Ideal R)] [DecidableEq <| Asso rw [← Ideal.dvd_iff_le, ← Associates.mk_dvd_mk, Associates.mk_pow] simp only [Associates.dvd_eq_le] rw [Associates.prime_pow_dvd_iff_le hI hJ'] - linarith + omega end diff --git a/Mathlib/RingTheory/LaurentSeries.lean b/Mathlib/RingTheory/LaurentSeries.lean index 1e09014fa6640..105b7f08e0505 100644 --- a/Mathlib/RingTheory/LaurentSeries.lean +++ b/Mathlib/RingTheory/LaurentSeries.lean @@ -633,7 +633,7 @@ theorem coeff_zero_of_lt_valuation {n D : ℤ} {f : K⸨X⸩} by_cases ord_nonpos : f.order ≤ 0 · obtain ⟨s, hs⟩ := Int.exists_eq_neg_ofNat ord_nonpos obtain ⟨m, hm⟩ := Int.eq_ofNat_of_zero_le (neg_le_iff_add_nonneg.mp (hs ▸ h_n_ord)) - obtain ⟨d, hd⟩ := Int.eq_ofNat_of_zero_le (a := D + s) (by linarith) + obtain ⟨d, hd⟩ := Int.eq_ofNat_of_zero_le (a := D + s) (by omega) rw [eq_add_neg_of_add_eq hm, add_comm, ← hs, ← powerSeriesPart_coeff] apply (intValuation_le_iff_coeff_lt_eq_zero K F).mp _ m (by linarith) rwa [hF, ofPowerSeries_powerSeriesPart f, hs, neg_neg, ← hd, neg_add_rev, ofAdd_add, map_mul, @@ -641,8 +641,8 @@ theorem coeff_zero_of_lt_valuation {n D : ℤ} {f : K⸨X⸩} mul_le_mul_left (by simp only [ne_eq, WithZero.coe_ne_zero, not_false_iff, zero_lt_iff])] · rw [not_le] at ord_nonpos obtain ⟨s, hs⟩ := Int.exists_eq_neg_ofNat (Int.neg_nonpos_of_nonneg (le_of_lt ord_nonpos)) - obtain ⟨m, hm⟩ := Int.eq_ofNat_of_zero_le (a := n - s) (by linarith) - obtain ⟨d, hd⟩ := Int.eq_ofNat_of_zero_le (a := D - s) (by linarith) + obtain ⟨m, hm⟩ := Int.eq_ofNat_of_zero_le (a := n - s) (by omega) + obtain ⟨d, hd⟩ := Int.eq_ofNat_of_zero_le (a := D - s) (by omega) rw [(sub_eq_iff_eq_add).mp hm, add_comm, ← neg_neg (s : ℤ), ← hs, neg_neg, ← powerSeriesPart_coeff] apply (intValuation_le_iff_coeff_lt_eq_zero K F).mp _ m (by linarith) @@ -670,7 +670,7 @@ theorem valuation_le_iff_coeff_lt_eq_zero {D : ℤ} {f : K⸨X⸩} : intro n hn rw [powerSeriesPart_coeff f n, hs] apply h_val_f - linarith + omega · simp only [ne_eq, WithZero.coe_ne_zero, not_false_iff, zero_lt_iff] · obtain ⟨s, hs⟩ := Int.exists_eq_neg_ofNat <| neg_nonpos_of_nonneg <| le_of_lt <| not_le.mp ord_nonpos @@ -681,14 +681,14 @@ theorem valuation_le_iff_coeff_lt_eq_zero {D : ℤ} {f : K⸨X⸩} : · by_cases hDs : D - s ≤ 0 · apply le_trans ((PowerSeries.idealX K).valuation_le_one F) rw [← WithZero.coe_one, ← ofAdd_zero, WithZero.coe_le_coe, Multiplicative.ofAdd_le] - linarith + omega · obtain ⟨d, hd⟩ := Int.eq_ofNat_of_zero_le (le_of_lt <| not_le.mp hDs) rw [← neg_neg (-D + ↑s), ← sub_eq_neg_add, neg_sub, hd] apply (intValuation_le_iff_coeff_lt_eq_zero K F).mpr intro n hn rw [powerSeriesPart_coeff f n, hs] apply h_val_f (s + n) - linarith + omega · simp only [ne_eq, WithZero.coe_ne_zero, not_false_iff, zero_lt_iff] /- Two Laurent series whose difference has small valuation have the same coefficients for @@ -717,7 +717,7 @@ theorem val_le_one_iff_eq_coe (f : K⸨X⸩) : Valued.v f ≤ (1 : ℤₘ₀) Set.mem_range, not_exists, Int.negSucc_lt_zero, reduceCtorEq] intro · simp only [not_false_eq_true] - · linarith + · omega end LaurentSeries diff --git a/Mathlib/RingTheory/PowerSeries/Basic.lean b/Mathlib/RingTheory/PowerSeries/Basic.lean index 2d18bb322885f..00d42f27bd55b 100644 --- a/Mathlib/RingTheory/PowerSeries/Basic.lean +++ b/Mathlib/RingTheory/PowerSeries/Basic.lean @@ -590,7 +590,7 @@ lemma coeff_one_pow (n : ℕ) (φ : R⟦X⟧) : rcases Nat.eq_zero_or_pos n with (rfl | hn) · simp induction n with - | zero => by_contra; linarith + | zero => by_contra; omega | succ n' ih => have h₁ (m : ℕ) : φ ^ (m + 1) = φ ^ m * φ := by exact rfl have h₂ : Finset.antidiagonal 1 = {(0, 1), (1, 0)} := by exact rfl diff --git a/Mathlib/RingTheory/RingHom/Integral.lean b/Mathlib/RingTheory/RingHom/Integral.lean index 2662ea9b1ccba..a938ac44a00cb 100644 --- a/Mathlib/RingTheory/RingHom/Integral.lean +++ b/Mathlib/RingTheory/RingHom/Integral.lean @@ -57,7 +57,7 @@ theorem isIntegral_ofLocalizationSpan : IsLocalization.map_eq_zero_iff (.powers (f t))] at hp' obtain ⟨⟨x, m, (rfl : algebraMap R S t ^ m = x)⟩, e⟩ := hp' by_cases hp' : 1 ≤ p.natDegree; swap - · obtain rfl : p = 1 := eq_one_of_monic_natDegree_zero hp (by nlinarith) + · obtain rfl : p = 1 := eq_one_of_monic_natDegree_zero hp (by omega) exact ⟨m, by simp [Algebra.smul_def, show algebraMap R S t ^ m = 0 by simpa using e]⟩ refine ⟨m + n, p.scaleRoots (t ^ m), (monic_scaleRoots_iff _).mpr hp, ?_⟩ have := p.scaleRoots_eval₂_mul (algebraMap R S) (t ^ n • r) (t ^ m) From d19d8405b0688fc5ac60e53056a822f6e53a34ec Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= <37772949+joelriou@users.noreply.github.com> Date: Sat, 14 Dec 2024 10:55:32 +0000 Subject: [PATCH 709/829] feat(CategoryTheory/Monoidal): composition of monoidal category adjunctions (#17956) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit If `C` and `D` are monoidal categories, `F : C ⥤ D` and `G : D ⥤ E` are monoidal functors, an adjunction `F ⊣ G` is monoidal if some additional compatibilities are satisfied. We show in this PR that the property that these compatibilities hold is stable by composition of adjunctions. Co-authored-by: Joël Riou --- Mathlib/CategoryTheory/Adjunction/Basic.lean | 14 +++++ Mathlib/CategoryTheory/Monoidal/Functor.lean | 55 +++++++++++++++++--- 2 files changed, 62 insertions(+), 7 deletions(-) diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean index 7bea78745ab75..9ba41bce3a283 100644 --- a/Mathlib/CategoryTheory/Adjunction/Basic.lean +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -179,6 +179,15 @@ attribute [local simp] Adjunction.homEquiv_unit Adjunction.homEquiv_counit namespace Adjunction +@[ext] +lemma ext {F : C ⥤ D} {G : D ⥤ C} {adj adj' : F ⊣ G} + (h : adj.unit = adj'.unit) : adj = adj' := by + suffices h' : adj.counit = adj'.counit by cases adj; cases adj'; aesop + ext X + apply (adj.homEquiv _ _).injective + rw [Adjunction.homEquiv_unit, Adjunction.homEquiv_unit, + Adjunction.right_triangle_components, h, Adjunction.right_triangle_components] + section variable {F : C ⥤ D} {G : D ⥤ C} (adj : F ⊣ G) @@ -659,6 +668,11 @@ lemma isLeftAdjoint_inverse : e.inverse.IsLeftAdjoint := lemma isRightAdjoint_functor : e.functor.IsRightAdjoint := e.symm.isRightAdjoint_inverse +lemma trans_toAdjunction {E : Type*} [Category E] (e' : D ≌ E) : + (e.trans e').toAdjunction = e.toAdjunction.comp e'.toAdjunction := by + ext + simp [trans] + end Equivalence namespace Functor diff --git a/Mathlib/CategoryTheory/Monoidal/Functor.lean b/Mathlib/CategoryTheory/Monoidal/Functor.lean index 83192bd9443b1..75e6de20c0015 100644 --- a/Mathlib/CategoryTheory/Monoidal/Functor.lean +++ b/Mathlib/CategoryTheory/Monoidal/Functor.lean @@ -959,6 +959,27 @@ lemma map_μ_comp_counit_app_tensor (X Y : D) : rw [IsMonoidal.leftAdjoint_μ (adj := adj), homEquiv_unit] simp +instance : (Adjunction.id (C := C)).IsMonoidal where + leftAdjoint_ε := by simp [id, homEquiv] + leftAdjoint_μ := by simp [id, homEquiv] + +instance isMonoidal_comp {F' : D ⥤ E} {G' : E ⥤ D} (adj' : F' ⊣ G') + [F'.OplaxMonoidal] [G'.LaxMonoidal] [adj'.IsMonoidal] : (adj.comp adj').IsMonoidal where + leftAdjoint_ε := by + dsimp [homEquiv] + rw [← adj.unit_app_unit_comp_map_η, ← adj'.unit_app_unit_comp_map_η, + assoc, comp_unit_app, assoc, ← Functor.map_comp, + ← adj'.unit_naturality_assoc, ← map_comp, ← map_comp] + leftAdjoint_μ X Y := by + apply ((adj.comp adj').homEquiv _ _).symm.injective + erw [Equiv.symm_apply_apply] + dsimp [homEquiv] + rw [comp_counit_app, comp_counit_app, comp_counit_app, assoc, tensor_comp, δ_natural_assoc] + dsimp + rw [← adj'.map_μ_comp_counit_app_tensor, ← map_comp_assoc, ← map_comp_assoc, + ← map_comp_assoc, ← adj.map_μ_comp_counit_app_tensor, assoc, + F.map_comp_assoc, counit_naturality] + end Adjunction namespace Equivalence @@ -980,7 +1001,7 @@ noncomputable def inverseMonoidal [e.functor.Monoidal] : e.inverse.Monoidal := b apply Monoidal.ofLaxMonoidal /-- An equivalence of categories involving monoidal functors is monoidal if the underlying -adjunction satisfies certain compatibilities with respect to the monoidal funtor data. -/ +adjunction satisfies certain compatibilities with respect to the monoidal functor data. -/ abbrev IsMonoidal [e.functor.Monoidal] [e.inverse.Monoidal] : Prop := e.toAdjunction.IsMonoidal example [e.functor.Monoidal] : letI := e.inverseMonoidal; e.IsMonoidal := inferInstance @@ -1011,12 +1032,6 @@ lemma functor_map_μ_inverse_comp_counitIso_hom_app_tensor (X Y : D) : δ e.functor _ _ ≫ (e.counitIso.hom.app X ⊗ e.counitIso.hom.app Y) := e.toAdjunction.map_μ_comp_counit_app_tensor X Y -@[reassoc] -lemma unitIso_hom_app_tensor_comp_inverse_map_δ_functor__ (X Y : C) : - e.unitIso.hom.app (X ⊗ Y) ≫ e.inverse.map (δ e.functor X Y) = - (e.unitIso.hom.app X ⊗ e.unitIso.hom.app Y) ≫ μ e.inverse _ _ := - e.toAdjunction.unit_app_tensor_comp_map_δ X Y - @[reassoc] lemma counitIso_inv_app_comp_functor_map_η_inverse : e.counitIso.inv.app (𝟙_ D) ≫ e.functor.map (η e.inverse) = ε e.functor := by @@ -1033,6 +1048,13 @@ lemma counitIso_inv_app_tensor_comp_functor_map_δ_inverse (X Y : C) : simp [← cancel_epi (e.unitIso.hom.app (X ⊗ Y)), Functor.map_comp, unitIso_hom_app_tensor_comp_inverse_map_δ_functor_assoc] +instance : (refl (C := C)).functor.Monoidal := inferInstanceAs (𝟭 C).Monoidal +instance : (refl (C := C)).inverse.Monoidal := inferInstanceAs (𝟭 C).Monoidal + +/-- The obvious auto-equivalence of a monoidal category is monoidal. -/ +instance isMonoidal_refl : (Equivalence.refl (C := C)).IsMonoidal := + inferInstanceAs (Adjunction.id (C := C)).IsMonoidal + /-- The inverse of a monoidal category equivalence is also a monoidal category equivalence. -/ instance isMonoidal_symm [e.inverse.Monoidal] [e.IsMonoidal] : e.symm.IsMonoidal where @@ -1048,6 +1070,25 @@ instance isMonoidal_symm [e.inverse.Monoidal] [e.IsMonoidal] : dsimp rw [tensorHom_id, id_whiskerRight, map_id, comp_id] +section + +variable (e' : D ≌ E) + +instance [e'.functor.Monoidal] : (e.trans e').functor.Monoidal := + inferInstanceAs (e.functor ⋙ e'.functor).Monoidal + +instance [e'.inverse.Monoidal] : (e.trans e').inverse.Monoidal := + inferInstanceAs (e'.inverse ⋙ e.inverse).Monoidal + +/-- The composition of two monoidal category equivalences is monoidal. -/ +instance isMonoidal_trans [e'.functor.Monoidal] [e'.inverse.Monoidal] [e'.IsMonoidal] : + (e.trans e').IsMonoidal := by + dsimp [Equivalence.IsMonoidal] + rw [trans_toAdjunction] + infer_instance + +end + end Equivalence variable (C D) From b09464fc7b0ff4bcfd4de7ff54289799009b5913 Mon Sep 17 00:00:00 2001 From: smorel394 Date: Sat, 14 Dec 2024 12:11:17 +0000 Subject: [PATCH 710/829] style(Algebra/Category/Grp/Basic and Algebra/Category/MonCat/Basic): universe order in `uliftFunctor` (#19957) Change the order of universes in the `uliftFunctor`s on the group and monoid categories so they will agree with the order used for `CategoryTheory.uliftFunctor` and `SSet.uliftFunctor` (that is, `uliftFunctor.{u,v}` should go from the category of doodads in `Type v` to the category of similar doodads in `Type (max v u)`). --- Mathlib/Algebra/Category/Grp/Basic.lean | 8 ++++---- Mathlib/Algebra/Category/MonCat/Basic.lean | 8 ++++---- 2 files changed, 8 insertions(+), 8 deletions(-) diff --git a/Mathlib/Algebra/Category/Grp/Basic.lean b/Mathlib/Algebra/Category/Grp/Basic.lean index 29dfb8506059b..7be93501da135 100644 --- a/Mathlib/Algebra/Category/Grp/Basic.lean +++ b/Mathlib/Algebra/Category/Grp/Basic.lean @@ -142,8 +142,8 @@ example {R S : Grp} (i : R ⟶ S) (r : R) (h : r = 1) : i r = 1 := by simp [h] /-- Universe lift functor for groups. -/ @[to_additive (attr := simps) "Universe lift functor for additive groups."] -def uliftFunctor : Grp.{u} ⥤ Grp.{max u v} where - obj X := Grp.of (ULift.{v, u} X) +def uliftFunctor : Grp.{v} ⥤ Grp.{max v u} where + obj X := Grp.of (ULift.{u, v} X) map {_ _} f := Grp.ofHom <| MulEquiv.ulift.symm.toMonoidHom.comp <| f.comp MulEquiv.ulift.toMonoidHom map_id X := by rfl @@ -283,8 +283,8 @@ example {R S : CommGrp} (i : R ⟶ S) (r : R) (h : r = 1) : i r = 1 := by simp [ /-- Universe lift functor for commutative groups. -/ @[to_additive (attr := simps) "Universe lift functor for additive commutative groups."] -def uliftFunctor : CommGrp.{u} ⥤ CommGrp.{max u v} where - obj X := CommGrp.of (ULift.{v, u} X) +def uliftFunctor : CommGrp.{v} ⥤ CommGrp.{max v u} where + obj X := CommGrp.of (ULift.{u, v} X) map {_ _} f := CommGrp.ofHom <| MulEquiv.ulift.symm.toMonoidHom.comp <| f.comp MulEquiv.ulift.toMonoidHom map_id X := by rfl diff --git a/Mathlib/Algebra/Category/MonCat/Basic.lean b/Mathlib/Algebra/Category/MonCat/Basic.lean index 981888898e4d7..7f1ee0139c0fc 100644 --- a/Mathlib/Algebra/Category/MonCat/Basic.lean +++ b/Mathlib/Algebra/Category/MonCat/Basic.lean @@ -143,8 +143,8 @@ instance {G : Type*} [Group G] : Group (MonCat.of G) := by assumption /-- Universe lift functor for monoids. -/ @[to_additive (attr := simps) "Universe lift functor for additive monoids."] -def uliftFunctor : MonCat.{u} ⥤ MonCat.{max u v} where - obj X := MonCat.of (ULift.{v, u} X) +def uliftFunctor : MonCat.{v} ⥤ MonCat.{max v u} where + obj X := MonCat.of (ULift.{u, v} X) map {_ _} f := MonCat.ofHom <| MulEquiv.ulift.symm.toMonoidHom.comp <| f.comp MulEquiv.ulift.toMonoidHom map_id X := by rfl @@ -247,8 +247,8 @@ lemma ofHom_apply {X Y : Type u} [CommMonoid X] [CommMonoid Y] (f : X →* Y) (x /-- Universe lift functor for commutative monoids. -/ @[to_additive (attr := simps) "Universe lift functor for additive commutative monoids."] -def uliftFunctor : CommMonCat.{u} ⥤ CommMonCat.{max u v} where - obj X := CommMonCat.of (ULift.{v, u} X) +def uliftFunctor : CommMonCat.{v} ⥤ CommMonCat.{max v u} where + obj X := CommMonCat.of (ULift.{u, v} X) map {_ _} f := CommMonCat.ofHom <| MulEquiv.ulift.symm.toMonoidHom.comp <| f.comp MulEquiv.ulift.toMonoidHom map_id X := by rfl From 468b507def7650050a687234d2cdf41a4d869429 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Sat, 14 Dec 2024 13:50:28 +0000 Subject: [PATCH 711/829] feat: drop an assumption in `mfderivWithin_const` (#19936) Followup to #19694 --- Mathlib/Geometry/Manifold/MFDeriv/Basic.lean | 9 ++++++++- Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean | 4 ++-- 2 files changed, 10 insertions(+), 3 deletions(-) diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean index acf6be63394ec..bd60a56508a1f 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean @@ -693,11 +693,18 @@ protected theorem MDifferentiableAt.mfderiv (h : MDifferentiableAt I I' f x) : protected theorem HasMFDerivAt.mfderiv (h : HasMFDerivAt I I' f x f') : mfderiv I I' f x = f' := (hasMFDerivAt_unique h h.mdifferentiableAt.hasMFDerivAt).symm -theorem HasMFDerivWithinAt.mfderivWithin (h : HasMFDerivWithinAt I I' f s x f') +protected theorem HasMFDerivWithinAt.mfderivWithin (h : HasMFDerivWithinAt I I' f s x f') (hxs : UniqueMDiffWithinAt I s x) : mfderivWithin I I' f s x = f' := by ext rw [hxs.eq h h.mdifferentiableWithinAt.hasMFDerivWithinAt] +theorem HasMFDerivWithinAt.mfderivWithin_eq_zero (h : HasMFDerivWithinAt I I' f s x 0) : + mfderivWithin I I' f s x = 0 := by + simp only [mfld_simps, mfderivWithin, h.mdifferentiableWithinAt, ↓reduceIte] + simp only [HasMFDerivWithinAt, mfld_simps] at h + rw [fderivWithin, if_pos] + exact h.2 + theorem MDifferentiable.mfderivWithin (h : MDifferentiableAt I I' f x) (hxs : UniqueMDiffWithinAt I s x) : mfderivWithin I I' f s x = mfderiv I I' f x := by apply HasMFDerivWithinAt.mfderivWithin _ hxs diff --git a/Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean b/Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean index 63e5592d2c537..8743fd3f8996c 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean @@ -198,9 +198,9 @@ theorem mfderiv_const : mfderiv I I' (fun _ : M => c) x = (0 : TangentSpace I x →L[𝕜] TangentSpace I' c) := HasMFDerivAt.mfderiv (hasMFDerivAt_const c x) -theorem mfderivWithin_const (hxs : UniqueMDiffWithinAt I s x) : +theorem mfderivWithin_const : mfderivWithin I I' (fun _ : M => c) s x = (0 : TangentSpace I x →L[𝕜] TangentSpace I' c) := - (hasMFDerivWithinAt_const _ _ _).mfderivWithin hxs + (hasMFDerivWithinAt_const _ _ _).mfderivWithin_eq_zero end Const From 79d29f84f37ec475db18dc902ccfdffab51ff0c6 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Sat, 14 Dec 2024 14:31:05 +0000 Subject: [PATCH 712/829] chore(discover-lean-pr-testing): only consider branches with "lean-pr-testing" in the name (#19958) Co-authored-by: adomani --- .../workflows/discover-lean-pr-testing.yml | 22 +++++++++---------- 1 file changed, 11 insertions(+), 11 deletions(-) diff --git a/.github/workflows/discover-lean-pr-testing.yml b/.github/workflows/discover-lean-pr-testing.yml index 0abbbef132b36..71d4a37aef33d 100644 --- a/.github/workflows/discover-lean-pr-testing.yml +++ b/.github/workflows/discover-lean-pr-testing.yml @@ -47,13 +47,13 @@ jobs: id: get-prs run: | NIGHTLY_URL="https://github.com/leanprover/lean4-nightly.git" - + # Create a temporary directory for cloning cd "$(mktemp -d)" || exit 1 - + # Clone the repository with a depth of 1 git clone --depth 1 "$NIGHTLY_URL" - + # Navigate to the cloned repository cd lean4-nightly || exit 1 @@ -71,7 +71,7 @@ jobs: # and extract the github PR numbers # (drop anything that doesn't look like a PR number from the final result) PRS=$(git log --oneline "$OLD..$NEW" | sed 's/.*(#\([0-9]\+\))$/\1/' | grep -E '^[0-9]+$') - + # Output the PRs echo "$PRS" printf "prs<> "$GITHUB_OUTPUT" @@ -86,25 +86,27 @@ jobs: echo "$PRS" | tr ' ' '\n' > prs.txt echo "=== prs.txt =====================" cat prs.txt - MATCHING_BRANCHES=$(git branch -r | grep -f prs.txt) + MATCHING_BRANCHES=$(git branch -r | grep -f prs.txt | grep "lean-pr-testing") echo "=== MATCHING_BRANCHES ===========" echo "$MATCHING_BRANCHES" + echo "=================================" # Initialize an empty variable to store branches with relevant diffs RELEVANT_BRANCHES="" # Loop through each matching branch for BRANCH in $MATCHING_BRANCHES; do + echo " === Testing $BRANCH for relevance." # Get the diff filenames DIFF_FILES=$(git diff --name-only "origin/nightly-testing...$BRANCH") - + # Check if the diff contains files other than the specified ones if echo "$DIFF_FILES" | grep -v -e 'lake-manifest.json' -e 'lakefile.lean' -e 'lean-toolchain'; then # Extract the PR number and actual branch name PR_NUMBER=$(echo "$BRANCH" | grep -oP '\d+$') ACTUAL_BRANCH=${BRANCH#origin/} GITHUB_DIFF="https://github.com/leanprover-community/mathlib4/compare/nightly-testing...$ACTUAL_BRANCH" - + # Append the branch details to RELEVANT_BRANCHES RELEVANT_BRANCHES="$RELEVANT_BRANCHES"$'\n- '"$ACTUAL_BRANCH (lean4#$PR_NUMBER) ([diff with \`nightly-testing\`]($GITHUB_DIFF))" fi @@ -115,10 +117,8 @@ jobs: echo "'$RELEVANT_BRANCHES'" printf "branches<> "$GITHUB_OUTPUT" - - name: Check if there are relevant branches - id: check-branches - run: | - if [ -z "${{ steps.find-branches.outputs.branches }}" ]; then + # Check if there are relevant branches, recording the outcome in the `branches_exist` variable + if [ -z "${RELEVANT_BRANCHES}" ]; then echo "branches_exist=false" >> "$GITHUB_ENV" else echo "branches_exist=true" >> "$GITHUB_ENV" From b8d062106bb776fdd5cb431bffe89ce6869cf547 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sat, 14 Dec 2024 15:31:12 +0000 Subject: [PATCH 713/829] feat(SetTheory/Cardinal/Cofinality): generalize `cof_preOmega` (#19916) This lemma is also true for `o = 0`, so we can take a `IsSuccPrelimit` argument instead of an `IsLimit` argument. --- Mathlib/SetTheory/Cardinal/Cofinality.lean | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index 786a594aa7434..5668b0eb5d566 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -664,8 +664,10 @@ theorem aleph0_le_cof {o} : ℵ₀ ≤ cof o ↔ IsLimit o := by exact not_succ_isLimit _ l @[simp] -theorem cof_preOmega {o : Ordinal} (ho : o.IsLimit) : (preOmega o).cof = o.cof := - isNormal_preOmega.cof_eq ho +theorem cof_preOmega {o : Ordinal} (ho : IsSuccPrelimit o) : (preOmega o).cof = o.cof := by + by_cases h : IsMin o + · simp [h.eq_bot] + · exact isNormal_preOmega.cof_eq ⟨h, ho⟩ @[simp] theorem cof_omega {o : Ordinal} (ho : o.IsLimit) : (ω_ o).cof = o.cof := From e2b32447acd18744718bc83f660bbce7856dfd85 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 14 Dec 2024 15:50:56 +0000 Subject: [PATCH 714/829] chore: make some `induction_on*` compatible with `induction` (#19898) --- .../Constructions/BorelSpace/Basic.lean | 24 +++---- .../Function/AEEqOfIntegral.lean | 9 ++- Mathlib/MeasureTheory/Group/Measure.lean | 11 ++- .../MeasureTheory/MeasurableSpace/Prod.lean | 35 ++++------ Mathlib/MeasureTheory/Measure/GiryMonad.lean | 9 ++- .../Measure/MeasureSpaceDef.lean | 3 +- Mathlib/MeasureTheory/Measure/Prod.lean | 31 ++++----- Mathlib/MeasureTheory/Measure/Restrict.lean | 17 ++--- .../MeasureTheory/Measure/Typeclasses.lean | 16 ++--- Mathlib/MeasureTheory/PiSystem.lean | 44 +++++++----- Mathlib/Probability/Independence/Kernel.lean | 67 ++++++++----------- .../Kernel/Disintegration/CDFToKernel.lean | 36 +++++----- .../Kernel/Disintegration/Density.lean | 47 +++++++------ .../Kernel/MeasurableIntegral.lean | 54 +++++---------- .../Probability/Martingale/Convergence.lean | 26 +++---- Mathlib/Topology/Baire/BaireMeasurable.lean | 34 +++++----- 16 files changed, 218 insertions(+), 245 deletions(-) diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean index 23adc61b8b596..3648b47ac5b12 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean @@ -210,16 +210,6 @@ instance (priority := 100) BorelSpace.countablyGenerated {α : Type*} [Topologic borelize α exact hb.borel_eq_generateFrom -theorem MeasurableSet.induction_on_open [TopologicalSpace α] [MeasurableSpace α] [BorelSpace α] - {C : Set α → Prop} (h_open : ∀ U, IsOpen U → C U) - (h_compl : ∀ t, MeasurableSet t → C t → C tᶜ) - (h_union : - ∀ f : ℕ → Set α, - Pairwise (Disjoint on f) → (∀ i, MeasurableSet (f i)) → (∀ i, C (f i)) → C (⋃ i, f i)) : - ∀ ⦃t⦄, MeasurableSet t → C t := - MeasurableSpace.induction_on_inter BorelSpace.measurable_eq isPiSystem_isOpen - (h_open _ isOpen_empty) h_open h_compl h_union - section variable [TopologicalSpace α] [MeasurableSpace α] [OpensMeasurableSpace α] [TopologicalSpace β] @@ -232,6 +222,16 @@ theorem IsOpen.measurableSet (h : IsOpen s) : MeasurableSet s := theorem IsOpen.nullMeasurableSet {μ} (h : IsOpen s) : NullMeasurableSet s μ := h.measurableSet.nullMeasurableSet +@[elab_as_elim] +theorem MeasurableSet.induction_on_open {C : ∀ s : Set γ, MeasurableSet s → Prop} + (isOpen : ∀ U (hU : IsOpen U), C U hU.measurableSet) + (compl : ∀ t (ht : MeasurableSet t), C t ht → C tᶜ ht.compl) + (iUnion : ∀ f : ℕ → Set γ, Pairwise (Disjoint on f) → ∀ (hf : ∀ i, MeasurableSet (f i)), + (∀ i, C (f i) (hf i)) → C (⋃ i, f i) (.iUnion hf)) : + ∀ t (ht : MeasurableSet t), C t ht := fun t ht ↦ + MeasurableSpace.induction_on_inter BorelSpace.measurable_eq isPiSystem_isOpen + (isOpen _ isOpen_empty) isOpen compl iUnion t ht + instance (priority := 1000) {s : Set α} [h : HasCountableSeparatingOn α IsOpen s] : CountablySeparated s := by rw [CountablySeparated.subtype_iff] @@ -265,8 +265,8 @@ theorem IsCompact.nullMeasurableSet [T2Space α] {μ} (h : IsCompact s) : NullMe then they can't be separated by a Borel measurable set. -/ theorem Inseparable.mem_measurableSet_iff {x y : γ} (h : Inseparable x y) {s : Set γ} (hs : MeasurableSet s) : x ∈ s ↔ y ∈ s := - hs.induction_on_open (C := fun s ↦ (x ∈ s ↔ y ∈ s)) (fun _ ↦ h.mem_open_iff) (fun _ _ ↦ Iff.not) - fun _ _ _ h ↦ by simp [h] + MeasurableSet.induction_on_open (fun _ ↦ h.mem_open_iff) (fun _ _ ↦ Iff.not) + (fun _ _ _ h ↦ by simp [h]) s hs /-- If `K` is a compact set in an R₁ space and `s ⊇ K` is a Borel measurable superset, then `s` includes the closure of `K` as well. -/ diff --git a/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean b/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean index 10f195e4aacff..c6b139c016199 100644 --- a/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean +++ b/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean @@ -593,11 +593,10 @@ lemma ae_eq_zero_of_forall_setIntegral_isClosed_eq_zero {μ : Measure β} {f : have I : ∫ x, f x ∂μ = 0 := by rw [← setIntegral_univ]; exact h'f _ isClosed_univ simpa [ht, I] using integral_add_compl t_meas hf intro s hs - refine MeasurableSet.induction_on_open (fun U hU ↦ ?_) A (fun g g_disj g_meas hg ↦ ?_) hs - · rw [← compl_compl U] - exact A _ hU.measurableSet.compl (h'f _ hU.isClosed_compl) - · rw [integral_iUnion g_meas g_disj hf.integrableOn] - simp [hg] + induction s, hs using MeasurableSet.induction_on_open with + | isOpen U hU => exact compl_compl U ▸ A _ hU.measurableSet.compl (h'f _ hU.isClosed_compl) + | compl s hs ihs => exact A s hs ihs + | iUnion g g_disj g_meas hg => simp [integral_iUnion g_meas g_disj hf.integrableOn, hg] @[deprecated (since := "2024-04-17")] alias ae_eq_zero_of_forall_set_integral_isClosed_eq_zero := diff --git a/Mathlib/MeasureTheory/Group/Measure.lean b/Mathlib/MeasureTheory/Group/Measure.lean index 449597022125a..066a497a06598 100644 --- a/Mathlib/MeasureTheory/Group/Measure.lean +++ b/Mathlib/MeasureTheory/Group/Measure.lean @@ -603,13 +603,10 @@ theorem measure_univ_of_isMulLeftInvariant [WeaklyLocallyCompactSpace G] [Noncom @[to_additive] lemma _root_.MeasurableSet.mul_closure_one_eq {s : Set G} (hs : MeasurableSet s) : s * (closure {1} : Set G) = s := by - apply MeasurableSet.induction_on_open (C := fun t ↦ t • (closure {1} : Set G) = t) ?_ ?_ ?_ hs - · intro U hU - exact hU.mul_closure_one_eq - · rintro t - ht - exact compl_mul_closure_one_eq_iff.2 ht - · rintro f - - h''f - simp only [iUnion_smul, h''f] + induction s, hs using MeasurableSet.induction_on_open with + | isOpen U hU => exact hU.mul_closure_one_eq + | compl t _ iht => exact compl_mul_closure_one_eq_iff.2 iht + | iUnion f _ _ ihf => simp_rw [iUnion_mul f, ihf] @[to_additive (attr := simp)] lemma measure_mul_closure_one (s : Set G) (μ : Measure G) : diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Prod.lean b/Mathlib/MeasureTheory/MeasurableSpace/Prod.lean index ce0ec4a3af7a0..68e38d6cac91f 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Prod.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Prod.lean @@ -73,29 +73,22 @@ lemma MeasurableEmbedding.prodMap {α β γ δ : Type*} {mα : MeasurableSpace {mβ : MeasurableSpace β} {mγ : MeasurableSpace γ} {mδ : MeasurableSpace δ} {f : α → β} {g : γ → δ} (hg : MeasurableEmbedding g) (hf : MeasurableEmbedding f) : MeasurableEmbedding (Prod.map g f) := by - have h_inj : Function.Injective fun x : γ × α => (g x.fst, f x.snd) := by - intro x y hxy - rw [← @Prod.mk.eta _ _ x, ← @Prod.mk.eta _ _ y] - simp only [Prod.mk.inj_iff] at hxy ⊢ - exact ⟨hg.injective hxy.1, hf.injective hxy.2⟩ - refine ⟨h_inj, ?_, ?_⟩ + refine ⟨hg.injective.prodMap hf.injective, ?_, ?_⟩ · exact (hg.measurable.comp measurable_fst).prod_mk (hf.measurable.comp measurable_snd) - · -- Induction using the π-system of rectangles - refine fun s hs => - @MeasurableSpace.induction_on_inter _ - (fun s => MeasurableSet ((fun x : γ × α => (g x.fst, f x.snd)) '' s)) _ _ - generateFrom_prod.symm isPiSystem_prod ?_ ?_ ?_ ?_ _ hs - · simp only [Set.image_empty, MeasurableSet.empty] - · rintro t ⟨t₁, ht₁, t₂, ht₂, rfl⟩ - rw [← Set.prod_image_image_eq] + · intro s hs + -- Induction using the π-system of rectangles + induction s, hs using induction_on_inter generateFrom_prod.symm isPiSystem_prod with + | empty => + simp only [Set.image_empty, MeasurableSet.empty] + | basic s hs => + obtain ⟨t₁, ht₁, t₂, ht₂, rfl⟩ := hs + simp_rw [Prod.map, ← prod_image_image_eq] exact (hg.measurableSet_image.mpr ht₁).prod (hf.measurableSet_image.mpr ht₂) - · intro t _ ht_m - rw [← Set.range_diff_image h_inj, ← Set.prod_range_range_eq] - exact - MeasurableSet.diff (MeasurableSet.prod hg.measurableSet_range hf.measurableSet_range) ht_m - · intro g _ _ hg - simp_rw [Set.image_iUnion] - exact MeasurableSet.iUnion hg + | compl s _ ihs => + rw [← range_diff_image (hg.injective.prodMap hf.injective), range_prod_map] + exact .diff (.prod hg.measurableSet_range hf.measurableSet_range) ihs + | iUnion f _ _ ihf => + simpa only [image_iUnion] using .iUnion ihf @[deprecated (since := "2024-12-11")] alias MeasurableEmbedding.prod_mk := MeasurableEmbedding.prodMap diff --git a/Mathlib/MeasureTheory/Measure/GiryMonad.lean b/Mathlib/MeasureTheory/Measure/GiryMonad.lean index fd5a97331a2ec..0a897efd74f54 100644 --- a/Mathlib/MeasureTheory/Measure/GiryMonad.lean +++ b/Mathlib/MeasureTheory/Measure/GiryMonad.lean @@ -69,11 +69,14 @@ theorem _root_.Measurable.measure_of_isPiSystem {μ : α → Measure β} [∀ a, (h_basic : ∀ s ∈ S, Measurable fun a ↦ μ a s) (h_univ : Measurable fun a ↦ μ a univ) : Measurable μ := by rw [measurable_measure] - refine MeasurableSpace.induction_on_inter hgen hpi (by simp) h_basic ?_ ?_ - · intro s hsm ihs + intro s hs + induction s, hs using MeasurableSpace.induction_on_inter hgen hpi with + | empty => simp + | basic s hs => exact h_basic s hs + | compl s hsm ihs => simp only [measure_compl hsm (measure_ne_top _ _)] exact h_univ.sub ihs - · intro f hfd hfm ihf + | iUnion f hfd hfm ihf => simpa only [measure_iUnion hfd hfm] using .ennreal_tsum ihf theorem _root_.Measurable.measure_of_isPiSystem_of_isProbabilityMeasure {μ : α → Measure β} diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean b/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean index a05138dd08fce..6abb79f9bd8ec 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean @@ -271,7 +271,8 @@ theorem _root_.MeasurableSpace.ae_induction_on_inter Pairwise (Disjoint on f) → (∀ i, MeasurableSet (f i)) → (∀ i, C x (f i)) → C x (⋃ i, f i)) : ∀ᵐ x ∂μ, ∀ ⦃t⦄, MeasurableSet t → C x t := by filter_upwards [h_empty, h_basic, h_compl, h_union] with x hx_empty hx_basic hx_compl hx_union - using MeasurableSpace.induction_on_inter h_eq h_inter hx_empty hx_basic hx_compl hx_union + using MeasurableSpace.induction_on_inter (C := fun t _ ↦ C x t) + h_eq h_inter hx_empty hx_basic hx_compl hx_union end ae diff --git a/Mathlib/MeasureTheory/Measure/Prod.lean b/Mathlib/MeasureTheory/Measure/Prod.lean index b1b1cf0cd1ea5..ef9fde4f6a463 100644 --- a/Mathlib/MeasureTheory/Measure/Prod.lean +++ b/Mathlib/MeasureTheory/Measure/Prod.lean @@ -70,23 +70,20 @@ variable {μ μ' : Measure α} {ν ν' : Measure β} {τ : Measure γ} a measurable function. `measurable_measure_prod_mk_left` is strictly more general. -/ theorem measurable_measure_prod_mk_left_finite [IsFiniteMeasure ν] {s : Set (α × β)} (hs : MeasurableSet s) : Measurable fun x => ν (Prod.mk x ⁻¹' s) := by - classical - refine induction_on_inter (C := fun s => Measurable fun x => ν (Prod.mk x ⁻¹' s)) - generateFrom_prod.symm isPiSystem_prod ?_ ?_ ?_ ?_ hs - · simp - · rintro _ ⟨s, hs, t, _, rfl⟩ - simp only [mk_preimage_prod_right_eq_if, measure_if] - exact measurable_const.indicator hs - · intro t ht h2t - simp_rw [preimage_compl, measure_compl (measurable_prod_mk_left ht) (measure_ne_top ν _)] - exact h2t.const_sub _ - · intro f h1f h2f h3f - simp_rw [preimage_iUnion] - have : ∀ b, ν (⋃ i, Prod.mk b ⁻¹' f i) = ∑' i, ν (Prod.mk b ⁻¹' f i) := fun b => - measure_iUnion (fun i j hij => Disjoint.preimage _ (h1f hij)) fun i => - measurable_prod_mk_left (h2f i) - simp_rw [this] - apply Measurable.ennreal_tsum h3f + induction s, hs using induction_on_inter generateFrom_prod.symm isPiSystem_prod with + | empty => simp + | basic s hs => + obtain ⟨s, hs, t, -, rfl⟩ := hs + classical simpa only [mk_preimage_prod_right_eq_if, measure_if] + using measurable_const.indicator hs + | compl s hs ihs => + simp_rw [preimage_compl, measure_compl (measurable_prod_mk_left hs) (measure_ne_top ν _)] + exact ihs.const_sub _ + | iUnion f hfd hfm ihf => + have (a : α) : ν (Prod.mk a ⁻¹' ⋃ i, f i) = ∑' i, ν (Prod.mk a ⁻¹' f i) := by + rw [preimage_iUnion, measure_iUnion] + exacts [hfd.mono fun _ _ ↦ .preimage _, fun i ↦ measurable_prod_mk_left (hfm i)] + simpa only [this] using Measurable.ennreal_tsum ihf /-- If `ν` is an s-finite measure, and `s ⊆ α × β` is measurable, then `x ↦ ν { y | (x, y) ∈ s }` is a measurable function. -/ diff --git a/Mathlib/MeasureTheory/Measure/Restrict.lean b/Mathlib/MeasureTheory/Measure/Restrict.lean index b3823303c8980..71525b1a85059 100644 --- a/Mathlib/MeasureTheory/Measure/Restrict.lean +++ b/Mathlib/MeasureTheory/Measure/Restrict.lean @@ -423,17 +423,18 @@ theorem ext_of_generateFrom_of_cover {S T : Set (Set α)} (h_gen : ‹_› = gen refine ext_of_sUnion_eq_univ hc hU fun t ht => ?_ ext1 u hu simp only [restrict_apply hu] - refine induction_on_inter h_gen h_inter ?_ (ST_eq t ht) ?_ ?_ hu - · simp only [Set.empty_inter, measure_empty] - · intro v hv hvt + induction u, hu using induction_on_inter h_gen h_inter with + | empty => simp only [Set.empty_inter, measure_empty] + | basic u hu => exact ST_eq _ ht _ hu + | compl u hu ihu => have := T_eq t ht - rw [Set.inter_comm] at hvt ⊢ - rwa [← measure_inter_add_diff t hv, ← measure_inter_add_diff t hv, ← hvt, + rw [Set.inter_comm] at ihu ⊢ + rwa [← measure_inter_add_diff t hu, ← measure_inter_add_diff t hu, ← ihu, ENNReal.add_right_inj] at this exact ne_top_of_le_ne_top (htop t ht) (measure_mono Set.inter_subset_left) - · intro f hfd hfm h_eq - simp only [← restrict_apply (hfm _), ← restrict_apply (MeasurableSet.iUnion hfm)] at h_eq ⊢ - simp only [measure_iUnion hfd hfm, h_eq] + | iUnion f hfd hfm ihf => + simp only [← restrict_apply (hfm _), ← restrict_apply (MeasurableSet.iUnion hfm)] at ihf ⊢ + simp only [measure_iUnion hfd hfm, ihf] /-- Two measures are equal if they are equal on the π-system generating the σ-algebra, and they are both finite on an increasing spanning sequence of sets in the π-system. diff --git a/Mathlib/MeasureTheory/Measure/Typeclasses.lean b/Mathlib/MeasureTheory/Measure/Typeclasses.lean index f8d833559d3e9..71c507d192908 100644 --- a/Mathlib/MeasureTheory/Measure/Typeclasses.lean +++ b/Mathlib/MeasureTheory/Measure/Typeclasses.lean @@ -1331,14 +1331,14 @@ theorem ext_on_measurableSpace_of_generate_finite {α} (m₀ : MeasurableSpace constructor rw [← h_univ] apply IsFiniteMeasure.measure_univ_lt_top - refine induction_on_inter hA hC (by simp) hμν ?_ ?_ hs - · intro t h1t h2t - have h1t_ : @MeasurableSet α m₀ t := h _ h1t - rw [@measure_compl α m₀ μ t h1t_ (@measure_ne_top α m₀ μ _ t), - @measure_compl α m₀ ν t h1t_ (@measure_ne_top α m₀ ν _ t), h_univ, h2t] - · intro f h1f h2f h3f - have h2f_ : ∀ i : ℕ, @MeasurableSet α m₀ (f i) := fun i => h _ (h2f i) - simp [measure_iUnion, h1f, h3f, h2f_] + induction s, hs using induction_on_inter hA hC with + | empty => simp + | basic t ht => exact hμν t ht + | compl t htm iht => + rw [measure_compl (h t htm) (measure_ne_top _ _), measure_compl (h t htm) (measure_ne_top _ _), + iht, h_univ] + | iUnion f hfd hfm ihf => + simp [measure_iUnion, hfd, h _ (hfm _), ihf] /-- Two finite measures are equal if they are equal on the π-system generating the σ-algebra (and `univ`). -/ diff --git a/Mathlib/MeasureTheory/PiSystem.lean b/Mathlib/MeasureTheory/PiSystem.lean index ff2222db98d9f..89a46dada7446 100644 --- a/Mathlib/MeasureTheory/PiSystem.lean +++ b/Mathlib/MeasureTheory/PiSystem.lean @@ -664,26 +664,34 @@ theorem generateFrom_eq {s : Set (Set α)} (hs : IsPiSystem s) : end DynkinSystem -theorem induction_on_inter {C : Set α → Prop} {s : Set (Set α)} [m : MeasurableSpace α] - (h_eq : m = generateFrom s) (h_inter : IsPiSystem s) (h_empty : C ∅) (h_basic : ∀ t ∈ s, C t) - (h_compl : ∀ t, MeasurableSet t → C t → C tᶜ) - (h_union : - ∀ f : ℕ → Set α, - Pairwise (Disjoint on f) → (∀ i, MeasurableSet (f i)) → (∀ i, C (f i)) → C (⋃ i, f i)) : - ∀ ⦃t⦄, MeasurableSet t → C t := +/-- Induction principle for measurable sets. +If `s` is a π-system that generates the product `σ`-algebra on `α` +and a predicate `C` defined on measurable sets is true + +- on the empty set; +- on each set `t ∈ s`; +- on the complement of a measurable set that satisfies `C`; +- on the union of a sequence of pairwise disjoint measurable sets that satisfy `C`, + +then it is true on all measurable sets in `α`. -/ +@[elab_as_elim] +theorem induction_on_inter {m : MeasurableSpace α} {C : ∀ s : Set α, MeasurableSet s → Prop} + {s : Set (Set α)} (h_eq : m = generateFrom s) (h_inter : IsPiSystem s) + (empty : C ∅ .empty) (basic : ∀ t (ht : t ∈ s), C t <| h_eq ▸ .basic t ht) + (compl : ∀ t (htm : MeasurableSet t), C t htm → C tᶜ htm.compl) + (iUnion : ∀ (f : ℕ → Set α), Pairwise (Disjoint on f) → ∀ (hfm : ∀ i, MeasurableSet (f i)), + (∀ i, C (f i) (hfm i)) → C (⋃ i, f i) (.iUnion hfm)) : + ∀ t (ht : MeasurableSet t), C t ht := by have eq : MeasurableSet = DynkinSystem.GenerateHas s := by rw [h_eq, DynkinSystem.generateFrom_eq h_inter] rfl - fun t ht => - have : DynkinSystem.GenerateHas s t := by rwa [eq] at ht - this.recOn h_basic h_empty - (fun {t} ht => - h_compl t <| by - rw [eq] - exact ht) - fun {f} hf ht => - h_union f hf fun i => by - rw [eq] - exact ht _ + suffices ∀ t (ht : DynkinSystem.GenerateHas s t), C t (eq ▸ ht) from + fun t ht ↦ this t (eq ▸ ht) + intro t ht + induction ht with + | basic u hu => exact basic u hu + | empty => exact empty + | @compl u hu ihu => exact compl _ (eq ▸ hu) ihu + | @iUnion f hfd hf ihf => exact iUnion f hfd (eq ▸ hf) ihf end MeasurableSpace diff --git a/Mathlib/Probability/Independence/Kernel.lean b/Mathlib/Probability/Independence/Kernel.lean index 3f4c1971fcafe..27dcf76e54543 100644 --- a/Mathlib/Probability/Independence/Kernel.lean +++ b/Mathlib/Probability/Independence/Kernel.lean @@ -45,7 +45,7 @@ definitions in the particular case of the usual independence notion. * `ProbabilityTheory.Kernel.IndepSets.Indep`: variant with two π-systems. -/ -open MeasureTheory MeasurableSpace +open Set MeasureTheory MeasurableSpace open scoped MeasureTheory ENNReal @@ -479,32 +479,22 @@ theorem IndepSets.indep_aux {m₂ m : MeasurableSpace Ω} ∀ᵐ a ∂μ, κ a (t1 ∩ t2) = κ a t1 * κ a t2 := by rcases eq_zero_or_isMarkovKernel κ with rfl | h · simp - refine @induction_on_inter _ (fun t ↦ ∀ᵐ a ∂μ, κ a (t1 ∩ t) = κ a t1 * κ a t) _ - m₂ hpm2 hp2 ?_ ?_ ?_ ?_ t2 ht2m - · simp only [Set.inter_empty, measure_empty, mul_zero, eq_self_iff_true, - Filter.eventually_true] - · exact fun t ht_mem_p2 ↦ hyp t1 t ht1 ht_mem_p2 - · intros t ht h - filter_upwards [h] with a ha - have : t1 ∩ tᶜ = t1 \ (t1 ∩ t) := by - rw [Set.diff_self_inter, Set.diff_eq_compl_inter, Set.inter_comm] - rw [this, - measure_diff Set.inter_subset_left (ht1m.inter (h2 _ ht)).nullMeasurableSet - (measure_ne_top (κ a) _), - measure_compl (h2 _ ht) (measure_ne_top (κ a) t), measure_univ, - ENNReal.mul_sub (fun _ _ ↦ measure_ne_top (κ a) _), mul_one, ha] - · intros f hf_disj hf_meas h - rw [← ae_all_iff] at h - filter_upwards [h] with a ha - rw [Set.inter_iUnion, measure_iUnion] - · rw [measure_iUnion hf_disj (fun i ↦ h2 _ (hf_meas i))] - rw [← ENNReal.tsum_mul_left] - congr with i - rw [ha i] - · intros i j hij - rw [Function.onFun, Set.inter_comm t1, Set.inter_comm t1] - exact Disjoint.inter_left _ (Disjoint.inter_right _ (hf_disj hij)) - · exact fun i ↦ ht1m.inter (h2 _ (hf_meas i)) + induction t2, ht2m using induction_on_inter hpm2 hp2 with + | empty => simp + | basic u hu => exact hyp t1 u ht1 hu + | compl u hu ihu => + filter_upwards [ihu] with a ha + rw [← Set.diff_eq, ← Set.diff_self_inter, + measure_diff inter_subset_left (ht1m.inter (h2 _ hu)).nullMeasurableSet (measure_ne_top _ _), + ha, measure_compl (h2 _ hu) (measure_ne_top _ _), measure_univ, ENNReal.mul_sub, mul_one] + exact fun _ _ ↦ measure_ne_top _ _ + | iUnion f hfd hfm ihf => + rw [← ae_all_iff] at ihf + filter_upwards [ihf] with a ha + rw [inter_iUnion, measure_iUnion, measure_iUnion hfd fun i ↦ h2 _ (hfm i)] + · simp only [ENNReal.tsum_mul_left, ha] + · exact hfd.mono fun i j h ↦ (h.inter_left' _).inter_right' _ + · exact fun i ↦ .inter ht1m (h2 _ <| hfm i) /-- The measurable space structures generated by independent pi-systems are independent. -/ theorem IndepSets.indep {m1 m2 m : MeasurableSpace Ω} {κ : Kernel α Ω} {μ : Measure α} @@ -515,18 +505,15 @@ theorem IndepSets.indep {m1 m2 m : MeasurableSpace Ω} {κ : Kernel α Ω} {μ : rcases eq_zero_or_isMarkovKernel κ with rfl | h · simp intros t1 t2 ht1 ht2 - refine @induction_on_inter _ (fun t ↦ ∀ᵐ (a : α) ∂μ, κ a (t ∩ t2) = κ a t * κ a t2) _ m1 hpm1 hp1 - ?_ ?_ ?_ ?_ _ ht1 - · simp only [Set.empty_inter, measure_empty, zero_mul, eq_self_iff_true, - Filter.eventually_true] - · intros t ht_mem_p1 - have ht1 : MeasurableSet[m] t := by - refine h1 _ ?_ - rw [hpm1] - exact measurableSet_generateFrom ht_mem_p1 - exact IndepSets.indep_aux h2 hp2 hpm2 hyp ht_mem_p1 ht1 ht2 - · intros t ht h - filter_upwards [h] with a ha + induction t1, ht1 using induction_on_inter hpm1 hp1 with + | empty => + simp only [Set.empty_inter, measure_empty, zero_mul, eq_self_iff_true, Filter.eventually_true] + | basic t ht => + refine IndepSets.indep_aux h2 hp2 hpm2 hyp ht (h1 _ ?_) ht2 + rw [hpm1] + exact measurableSet_generateFrom ht + | compl t ht iht => + filter_upwards [iht] with a ha have : tᶜ ∩ t2 = t2 \ (t ∩ t2) := by rw [Set.inter_comm t, Set.diff_self_inter, Set.diff_eq_compl_inter] rw [this, Set.inter_comm t t2, @@ -534,7 +521,7 @@ theorem IndepSets.indep {m1 m2 m : MeasurableSpace Ω} {κ : Kernel α Ω} {μ : (measure_ne_top (κ a) _), Set.inter_comm, ha, measure_compl (h1 _ ht) (measure_ne_top (κ a) t), measure_univ, mul_comm (1 - κ a t), ENNReal.mul_sub (fun _ _ ↦ measure_ne_top (κ a) _), mul_one, mul_comm] - · intros f hf_disj hf_meas h + | iUnion f hf_disj hf_meas h => rw [← ae_all_iff] at h filter_upwards [h] with a ha rw [Set.inter_comm, Set.inter_iUnion, measure_iUnion] diff --git a/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean index b1fdd2d1f0025..3dceea49cb6c1 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean @@ -539,11 +539,13 @@ lemma setLIntegral_toKernel_prod [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ -- `setLIntegral_toKernel_Iic` gives the result for `t = Iic x`. These sets form a -- π-system that generates the Borel σ-algebra, hence we can get the same equality for any -- measurable set `t`. - apply MeasurableSpace.induction_on_inter (borel_eq_generateFrom_Iic ℝ) isPiSystem_Iic _ _ _ _ ht - · simp only [measure_empty, lintegral_const, zero_mul, prod_empty] - · rintro t ⟨q, rfl⟩ + induction t, ht + using MeasurableSpace.induction_on_inter (borel_eq_generateFrom_Iic ℝ) isPiSystem_Iic with + | empty => simp only [measure_empty, lintegral_const, zero_mul, prod_empty] + | basic t ht => + obtain ⟨q, rfl⟩ := ht exact setLIntegral_toKernel_Iic hf a _ hs - · intro t ht ht_lintegral + | compl t ht iht => calc ∫⁻ b in s, hf.toKernel f (a, b) tᶜ ∂(ν a) = ∫⁻ b in s, hf.toKernel f (a, b) univ - hf.toKernel f (a, b) t ∂(ν a) := by congr with x; rw [measure_compl ht (measure_ne_top (hf.toKernel f (a, x)) _)] @@ -551,24 +553,22 @@ lemma setLIntegral_toKernel_prod [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ - ∫⁻ b in s, hf.toKernel f (a, b) t ∂(ν a) := by rw [lintegral_sub] · exact (Kernel.measurable_coe (hf.toKernel f) ht).comp measurable_prod_mk_left - · rw [ht_lintegral] + · rw [iht] exact measure_ne_top _ _ · exact Eventually.of_forall fun a ↦ measure_mono (subset_univ _) _ = κ a (s ×ˢ univ) - κ a (s ×ˢ t) := by - rw [setLIntegral_toKernel_univ hf a hs, ht_lintegral] + rw [setLIntegral_toKernel_univ hf a hs, iht] _ = κ a (s ×ˢ tᶜ) := by rw [← measure_diff _ (hs.prod ht).nullMeasurableSet (measure_ne_top _ _)] · rw [prod_diff_prod, compl_eq_univ_diff] simp only [diff_self, empty_prod, union_empty] · rw [prod_subset_prod_iff] exact Or.inl ⟨subset_rfl, subset_univ t⟩ - · intro f hf_disj hf_meas hf_eq + | iUnion f hf_disj hf_meas ihf => simp_rw [measure_iUnion hf_disj hf_meas] rw [lintegral_tsum, prod_iUnion, measure_iUnion] - · simp_rw [hf_eq] - · intro i j hij - rw [Function.onFun, Set.disjoint_prod] - exact Or.inr (hf_disj hij) + · simp_rw [ihf] + · exact hf_disj.mono fun i j h ↦ h.set_prod_right _ _ · exact fun i ↦ MeasurableSet.prod hs (hf_meas i) · exact fun i ↦ ((Kernel.measurable_coe _ (hf_meas i)).comp measurable_prod_mk_left).aemeasurable.restrict @@ -582,10 +582,12 @@ lemma lintegral_toKernel_mem [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ ν) -- `setLIntegral_toKernel_prod` gives the result for sets of the form `t₁ × t₂`. These -- sets form a π-system that generates the product σ-algebra, hence we can get the same equality -- for any measurable set `s`. - apply MeasurableSpace.induction_on_inter generateFrom_prod.symm isPiSystem_prod _ _ _ _ hs - · simp only [mem_empty_iff_false, setOf_false, measure_empty, lintegral_const, - zero_mul] - · rintro _ ⟨t₁, ht₁, t₂, ht₂, rfl⟩ + induction s, hs + using MeasurableSpace.induction_on_inter generateFrom_prod.symm isPiSystem_prod with + | empty => + simp only [mem_empty_iff_false, setOf_false, measure_empty, lintegral_const, zero_mul] + | basic s hs => + rcases hs with ⟨t₁, ht₁, t₂, ht₂, rfl⟩ simp only [mem_setOf_eq] at ht₁ ht₂ have h_prod_eq_snd : ∀ a ∈ t₁, {x : ℝ | (a, x) ∈ t₁ ×ˢ t₂} = t₂ := by intro a ha @@ -606,7 +608,7 @@ lemma lintegral_toKernel_mem [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ ν) simp only [hat₁, prod_mk_mem_set_prod_eq, false_and, setOf_false, measure_empty] rw [h_eq1, h_eq2, add_zero] exact setLIntegral_toKernel_prod hf a ht₁ ht₂ - · intro t ht ht_eq + | compl t ht ht_eq => calc ∫⁻ b, hf.toKernel f (a, b) {y : ℝ | (b, y) ∈ tᶜ} ∂(ν a) = ∫⁻ b, hf.toKernel f (a, b) {y : ℝ | (b, y) ∈ t}ᶜ ∂(ν a) := rfl _ = ∫⁻ b, hf.toKernel f (a, b) univ @@ -626,7 +628,7 @@ lemma lintegral_toKernel_mem [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ ν) exact measure_lt_top _ univ _ = κ a univ - κ a t := by rw [ht_eq, lintegral_toKernel_univ hf] _ = κ a tᶜ := (measure_compl ht (measure_ne_top _ _)).symm - · intro f' hf_disj hf_meas hf_eq + | iUnion f' hf_disj hf_meas hf_eq => have h_eq : ∀ a, {x | (a, x) ∈ ⋃ i, f' i} = ⋃ i, {x | (a, x) ∈ f' i} := by intro a; ext x; simp only [mem_iUnion, mem_setOf_eq] simp_rw [h_eq] diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index 316c64c2fce93..30f8ae52d6d41 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -560,18 +560,21 @@ lemma setIntegral_density (hκν : fst κ ≤ ν) [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) {A : Set γ} (hA : MeasurableSet A) : ∫ x in A, density κ ν a x s ∂(ν a) = (κ a (A ×ˢ s)).toReal := by have : IsFiniteKernel κ := isFiniteKernel_of_isFiniteKernel_fst (h := isFiniteKernel_of_le hκν) - have hA' : MeasurableSet[⨆ n, countableFiltration γ n] A := by rwa [iSup_countableFiltration] - refine induction_on_inter (m := ⨆ n, countableFiltration γ n) - (C := fun A ↦ ∫ x in A, density κ ν a x s ∂(ν a) = (κ a (A ×ˢ s)).toReal) - (measurableSpace_iSup_eq (countableFiltration γ)) ?_ ?_ ?_ ?_ ?_ hA' - · rintro s ⟨n, hs⟩ t ⟨m, ht⟩ _ - exact ⟨max n m, ((countableFiltration γ).mono (le_max_left n m) _ hs).inter - ((countableFiltration γ).mono (le_max_right n m) _ ht)⟩ - · simp - · intro A ⟨n, hA⟩ - exact setIntegral_density_of_measurableSet hκν n a hs hA - · intro A hA hA_eq - rw [iSup_countableFiltration] at hA + have hgen : ‹MeasurableSpace γ› = + .generateFrom {s | ∃ n, MeasurableSet[countableFiltration γ n] s} := by + rw [setOf_exists, generateFrom_iUnion_measurableSet (countableFiltration γ), + iSup_countableFiltration] + have hpi : IsPiSystem {s | ∃ n, MeasurableSet[countableFiltration γ n] s} := by + rw [setOf_exists] + exact isPiSystem_iUnion_of_monotone _ + (fun n ↦ @isPiSystem_measurableSet _ (countableFiltration γ n)) + fun _ _ ↦ (countableFiltration γ).mono + induction A, hA using induction_on_inter hgen hpi with + | empty => simp + | basic s hs => + rcases hs with ⟨n, hn⟩ + exact setIntegral_density_of_measurableSet hκν n a hs hn + | compl A hA hA_eq => have h := integral_add_compl hA (integrable_density hκν a hs) rw [hA_eq, integral_density hκν a hs] at h have : Aᶜ ×ˢ s = univ ×ˢ s \ A ×ˢ s := by @@ -582,18 +585,14 @@ lemma setIntegral_density (hκν : fst κ ≤ ν) [IsFiniteKernel ν] rw [eq_tsub_iff_add_eq_of_le, add_comm] · exact h · gcongr <;> simp - · intro f hf_disj hf h_eq - rw [integral_iUnion _ hf_disj (integrable_density hκν _ hs).integrableOn] - · simp_rw [h_eq] - rw [← ENNReal.tsum_toReal_eq (fun _ ↦ measure_ne_top _ _)] - congr - rw [iUnion_prod_const, measure_iUnion] - · intro i j hij - rw [Function.onFun, Set.disjoint_prod] - exact Or.inl (hf_disj hij) - · rw [iSup_countableFiltration] at hf - exact fun i ↦ (hf i).prod hs - · rwa [iSup_countableFiltration] at hf + | iUnion f hf_disj hf h_eq => + rw [integral_iUnion hf hf_disj (integrable_density hκν _ hs).integrableOn] + simp_rw [h_eq] + rw [← ENNReal.tsum_toReal_eq (fun _ ↦ measure_ne_top _ _)] + congr + rw [iUnion_prod_const, measure_iUnion] + · exact hf_disj.mono fun _ _ h ↦ h.set_prod_left _ _ + · exact fun i ↦ (hf i).prod hs @[deprecated (since := "2024-04-17")] alias set_integral_density := setIntegral_density diff --git a/Mathlib/Probability/Kernel/MeasurableIntegral.lean b/Mathlib/Probability/Kernel/MeasurableIntegral.lean index f200e6bcd9153..7c6f7a6e6bd19 100644 --- a/Mathlib/Probability/Kernel/MeasurableIntegral.lean +++ b/Mathlib/Probability/Kernel/MeasurableIntegral.lean @@ -40,16 +40,13 @@ theorem measurable_kernel_prod_mk_left_of_finite {t : Set (α × β)} (ht : Meas (hκs : ∀ a, IsFiniteMeasure (κ a)) : Measurable fun a => κ a (Prod.mk a ⁻¹' t) := by -- `t` is a measurable set in the product `α × β`: we use that the product σ-algebra is generated -- by boxes to prove the result by induction. - -- Porting note: added motive - refine MeasurableSpace.induction_on_inter - (C := fun t => Measurable fun a => κ a (Prod.mk a ⁻¹' t)) - generateFrom_prod.symm isPiSystem_prod ?_ ?_ ?_ ?_ ht - ·-- case `t = ∅` + induction t, ht + using MeasurableSpace.induction_on_inter generateFrom_prod.symm isPiSystem_prod with + | empty => simp only [preimage_empty, measure_empty, measurable_const] - · -- case of a box: `t = t₁ ×ˢ t₂` for measurable sets `t₁` and `t₂` - intro t' ht' - simp only [Set.mem_image2, Set.mem_setOf_eq, exists_and_left] at ht' - obtain ⟨t₁, ht₁, t₂, ht₂, rfl⟩ := ht' + | basic t ht => + simp only [Set.mem_image2, Set.mem_setOf_eq] at ht + obtain ⟨t₁, ht₁, t₂, ht₂, rfl⟩ := ht classical simp_rw [mk_preimage_prod_right_eq_if] have h_eq_ite : (fun a => κ a (ite (a ∈ t₁) t₂ ∅)) = fun a => ite (a ∈ t₁) (κ a t₂) 0 := by @@ -58,42 +55,27 @@ theorem measurable_kernel_prod_mk_left_of_finite {t : Set (α × β)} (ht : Meas exacts [rfl, measure_empty] rw [h_eq_ite] exact Measurable.ite ht₁ (Kernel.measurable_coe κ ht₂) measurable_const - · -- we assume that the result is true for `t` and we prove it for `tᶜ` - intro t' ht' h_meas - have h_eq_sdiff : ∀ a, Prod.mk a ⁻¹' t'ᶜ = Set.univ \ Prod.mk a ⁻¹' t' := by + | compl t htm iht => + have h_eq_sdiff : ∀ a, Prod.mk a ⁻¹' tᶜ = Set.univ \ Prod.mk a ⁻¹' t := by intro a ext1 b simp only [mem_compl_iff, mem_preimage, mem_diff, mem_univ, true_and] simp_rw [h_eq_sdiff] have : - (fun a => κ a (Set.univ \ Prod.mk a ⁻¹' t')) = fun a => - κ a Set.univ - κ a (Prod.mk a ⁻¹' t') := by + (fun a => κ a (Set.univ \ Prod.mk a ⁻¹' t)) = fun a => + κ a Set.univ - κ a (Prod.mk a ⁻¹' t) := by ext1 a rw [← Set.diff_inter_self_eq_diff, Set.inter_univ, measure_diff (Set.subset_univ _)] - · exact (measurable_prod_mk_left ht').nullMeasurableSet + · exact (measurable_prod_mk_left htm).nullMeasurableSet · exact measure_ne_top _ _ rw [this] - exact Measurable.sub (Kernel.measurable_coe κ MeasurableSet.univ) h_meas - · -- we assume that the result is true for a family of disjoint sets and prove it for their union - intro f h_disj hf_meas hf - have h_Union : - (fun a => κ a (Prod.mk a ⁻¹' ⋃ i, f i)) = fun a => κ a (⋃ i, Prod.mk a ⁻¹' f i) := by - ext1 a - congr with b - simp only [mem_iUnion, mem_preimage] - rw [h_Union] - have h_tsum : - (fun a => κ a (⋃ i, Prod.mk a ⁻¹' f i)) = fun a => ∑' i, κ a (Prod.mk a ⁻¹' f i) := by - ext1 a - rw [measure_iUnion] - · intro i j hij s hsi hsj b hbs - have habi : {(a, b)} ⊆ f i := by rw [Set.singleton_subset_iff]; exact hsi hbs - have habj : {(a, b)} ⊆ f j := by rw [Set.singleton_subset_iff]; exact hsj hbs - simpa only [Set.bot_eq_empty, Set.le_eq_subset, Set.singleton_subset_iff, - Set.mem_empty_iff_false] using h_disj hij habi habj - · exact fun i => (@measurable_prod_mk_left α β _ _ a) (hf_meas i) - rw [h_tsum] - exact Measurable.ennreal_tsum hf + exact Measurable.sub (Kernel.measurable_coe κ MeasurableSet.univ) iht + | iUnion f h_disj hf_meas hf => + have (a : α) : κ a (Prod.mk a ⁻¹' ⋃ i, f i) = ∑' i, κ a (Prod.mk a ⁻¹' f i) := by + rw [preimage_iUnion, measure_iUnion] + · exact h_disj.mono fun _ _ ↦ .preimage _ + · exact fun i ↦ measurable_prod_mk_left (hf_meas i) + simpa only [this] using Measurable.ennreal_tsum hf theorem measurable_kernel_prod_mk_left [IsSFiniteKernel κ] {t : Set (α × β)} (ht : MeasurableSet t) : Measurable fun a => κ a (Prod.mk a ⁻¹' t) := by diff --git a/Mathlib/Probability/Martingale/Convergence.lean b/Mathlib/Probability/Martingale/Convergence.lean index cd3dced4e0672..dbea5566f7eab 100644 --- a/Mathlib/Probability/Martingale/Convergence.lean +++ b/Mathlib/Probability/Martingale/Convergence.lean @@ -379,19 +379,19 @@ theorem Integrable.tendsto_ae_condexp (hg : Integrable g μ) filter_upwards [(martingale_condexp g ℱ μ).ae_eq_condexp_limitProcess hunif n] with x hx _ rw [hx] refine ae_eq_of_forall_setIntegral_eq_of_sigmaFinite' hle (fun s _ _ => hg.integrableOn) - (fun s _ _ => hlimint.integrableOn) (fun s hs => ?_) hgmeas.aeStronglyMeasurable' + (fun s _ _ => hlimint.integrableOn) (fun s hs _ => ?_) hgmeas.aeStronglyMeasurable' stronglyMeasurable_limitProcess.aeStronglyMeasurable' - apply @MeasurableSpace.induction_on_inter _ _ _ (⨆ n, ℱ n) - (MeasurableSpace.measurableSpace_iSup_eq ℱ) _ _ _ _ _ _ hs - · rintro s ⟨n, hs⟩ t ⟨m, ht⟩ - - by_cases hnm : n ≤ m - · exact ⟨m, (ℱ.mono hnm _ hs).inter ht⟩ - · exact ⟨n, hs.inter (ℱ.mono (not_le.1 hnm).le _ ht)⟩ - · simp only [measure_empty, ENNReal.zero_lt_top, Measure.restrict_empty, integral_zero_measure, - forall_true_left] - · rintro t ⟨n, ht⟩ - - exact this n _ ht - · rintro t htmeas ht - + have hpi : IsPiSystem {s | ∃ n, MeasurableSet[ℱ n] s} := by + rw [Set.setOf_exists] + exact isPiSystem_iUnion_of_monotone _ (fun n ↦ (ℱ n).isPiSystem_measurableSet) fun _ _ ↦ ℱ.mono + induction s, hs + using MeasurableSpace.induction_on_inter (MeasurableSpace.measurableSpace_iSup_eq ℱ) hpi with + | empty => + simp only [measure_empty, Measure.restrict_empty, integral_zero_measure] + | basic s hs => + rcases hs with ⟨n, hn⟩ + exact this n _ hn + | compl t htmeas ht => have hgeq := @setIntegral_compl _ _ (⨆ n, ℱ n) _ _ _ _ _ htmeas (hg.trim hle hgmeas) have hheq := @setIntegral_compl _ _ (⨆ n, ℱ n) _ _ _ _ _ htmeas (hlimint.trim hle stronglyMeasurable_limitProcess) @@ -401,7 +401,7 @@ theorem Integrable.tendsto_ae_condexp (hg : Integrable g μ) setIntegral_trim hle stronglyMeasurable_limitProcess htmeas, ← integral_trim hle hgmeas, ← integral_trim hle stronglyMeasurable_limitProcess, ← setIntegral_univ, this 0 _ MeasurableSet.univ, setIntegral_univ, ht (measure_lt_top _ _)] - · rintro f hf hfmeas heq - + | iUnion f hf hfmeas heq => rw [integral_iUnion (fun n => hle _ (hfmeas n)) hf hg.integrableOn, integral_iUnion (fun n => hle _ (hfmeas n)) hf hlimint.integrableOn] exact tsum_congr fun n => heq _ (measure_lt_top _ _) diff --git a/Mathlib/Topology/Baire/BaireMeasurable.lean b/Mathlib/Topology/Baire/BaireMeasurable.lean index 1c08bafa7d149..86cceb1e9825f 100644 --- a/Mathlib/Topology/Baire/BaireMeasurable.lean +++ b/Mathlib/Topology/Baire/BaireMeasurable.lean @@ -43,6 +43,15 @@ scoped[Topology] notation3 "∃ᵇ "(...)", "r:(scoped p => Filter.Frequently p variable {α} +theorem coborder_mem_residual {s : Set α} (hs : IsLocallyClosed s) : coborder s ∈ residual α := + residual_of_dense_open hs.isOpen_coborder dense_coborder + +theorem closure_residualEq {s : Set α} (hs : IsLocallyClosed s) : closure s =ᵇ s := by + rw [Filter.eventuallyEq_set] + filter_upwards [coborder_mem_residual hs] with x hx + nth_rewrite 2 [← closure_inter_coborder (s := s)] + simp [hx] + /-- We say a set is a `BaireMeasurableSet` if it differs from some Borel set by a meager set. This forms a σ-algebra. @@ -119,21 +128,16 @@ open Filter /--Any Borel set differs from some open set by a meager set. -/ theorem MeasurableSet.residualEq_isOpen [MeasurableSpace α] [BorelSpace α] (h : MeasurableSet s) : - ∃ u : Set α, (IsOpen u) ∧ s =ᵇ u := by - apply h.induction_on_open (fun s hs => ⟨s, hs, EventuallyEq.rfl⟩) - · rintro s - ⟨u, uo, su⟩ - refine ⟨(closure u)ᶜ, isClosed_closure.isOpen_compl, - EventuallyEq.compl (su.trans <| EventuallyLE.antisymm subset_closure.eventuallyLE ?_)⟩ - have : (coborder u) ∈ residual _ := - residual_of_dense_open uo.isLocallyClosed.isOpen_coborder dense_coborder - rw [coborder_eq_union_closure_compl] at this - rw [EventuallyLE] - convert this - simp only [le_Prop_eq, imp_iff_or_not] - rfl --terrible - rintro s - - hs - choose u uo su using hs - exact ⟨⋃ i, u i, isOpen_iUnion uo, EventuallyEq.countable_iUnion su⟩ + ∃ u : Set α, IsOpen u ∧ s =ᵇ u := by + induction s, h using MeasurableSet.induction_on_open with + | isOpen U hU => exact ⟨U, hU, .rfl⟩ + | compl s _ ihs => + obtain ⟨U, Uo, hsU⟩ := ihs + use (closure U)ᶜ, isClosed_closure.isOpen_compl + exact .compl <| hsU.trans <| .symm <| closure_residualEq Uo.isLocallyClosed + | iUnion f _ _ ihf => + choose u uo su using ihf + exact ⟨⋃ i, u i, isOpen_iUnion uo, EventuallyEq.countable_iUnion su⟩ /--Any `BaireMeasurableSet` differs from some open set by a meager set. -/ theorem BaireMeasurableSet.residualEq_isOpen (h : BaireMeasurableSet s) : From 197eebfa5455aa09b7a4c2ad3c3eb9123245df90 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Sat, 14 Dec 2024 15:50:57 +0000 Subject: [PATCH 715/829] feat: volume of Euclidean balls in even and odd dimensions (#19949) The formula in terms of the Gamma function is nice, as it is not necessary to split on whether the dimension is even or odd, but for specific calculations (e.g., the volume of a ball in three dimensions), it's nice to have these more specific versions. Written as a consequence of [Zulip](https://leanprover.zulipchat.com/#narrow/channel/113489-new-members/topic/How.20can.20I.20convert.20a.20sphere.20into.20a.20ball.3F). --- .../Gaussian/GaussianIntegral.lean | 22 +++++ .../Measure/Lebesgue/VolumeOfBalls.lean | 87 ++++++++++++++++--- .../CanonicalEmbedding/ConvexBody.lean | 1 + 3 files changed, 97 insertions(+), 13 deletions(-) diff --git a/Mathlib/Analysis/SpecialFunctions/Gaussian/GaussianIntegral.lean b/Mathlib/Analysis/SpecialFunctions/Gaussian/GaussianIntegral.lean index b88d8363dabf7..812fe616e1ee9 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gaussian/GaussianIntegral.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gaussian/GaussianIntegral.lean @@ -6,6 +6,7 @@ Authors: Sébastien Gouëzel import Mathlib.Analysis.SpecialFunctions.Gamma.Basic import Mathlib.Analysis.SpecialFunctions.PolarCoord import Mathlib.Analysis.Complex.Convex +import Mathlib.Data.Nat.Factorial.DoubleFactorial /-! # Gaussian integral @@ -351,3 +352,24 @@ theorem Complex.Gamma_one_half_eq : Complex.Gamma (1 / 2) = (π : ℂ) ^ (1 / 2 convert congr_arg ((↑) : ℝ → ℂ) Real.Gamma_one_half_eq · simpa only [one_div, ofReal_inv, ofReal_ofNat] using Gamma_ofReal (1 / 2) · rw [sqrt_eq_rpow, ofReal_cpow pi_pos.le, ofReal_div, ofReal_ofNat, ofReal_one] + +open scoped Nat in +/-- The special-value formula `Γ(k + 1 + 1/2) = (2 * k + 1)‼ * √π / (2 ^ (k + 1))` for half-integer +values of the gamma function in terms of `Nat.doubleFactorial`. -/ +lemma Real.Gamma_nat_add_one_add_half (k : ℕ) : + Gamma (k + 1 + 1 / 2) = (2 * k + 1 : ℕ)‼ * √π / (2 ^ (k + 1)) := by + induction k with + | zero => simp [-one_div, add_comm (1 : ℝ), Gamma_add_one, Gamma_one_half_eq]; ring + | succ k ih => + rw [add_right_comm, Gamma_add_one (by positivity), Nat.cast_add, Nat.cast_one, ih] + field_simp + ring + +open scoped Nat in +/-- The special-value formula `Γ(k + 1/2) = (2 * k - 1)‼ * √π / (2 ^ k))` for half-integer +values of the gamma function in terms of `Nat.doubleFactorial`. -/ +lemma Real.Gamma_nat_add_half (k : ℕ) : + Gamma (k + 1 / 2) = (2 * k - 1 : ℕ)‼ * √π / (2 ^ k) := by + cases k with + | zero => simp [- one_div, Gamma_one_half_eq] + | succ k => simpa [-one_div, mul_add] using Gamma_nat_add_one_add_half k diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean index 90b9c3cc0537b..fce565dd079d5 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean @@ -3,11 +3,11 @@ Copyright (c) 2023 Xavier Roblot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Xavier Roblot -/ -import Mathlib.Analysis.SpecialFunctions.Gamma.BohrMollerup import Mathlib.Data.Complex.FiniteDimensional import Mathlib.MeasureTheory.Constructions.HaarToSphere import Mathlib.MeasureTheory.Integral.Gamma import Mathlib.MeasureTheory.Integral.Pi +import Mathlib.Analysis.SpecialFunctions.Gaussian.GaussianIntegral /-! # Volume of balls @@ -314,13 +314,13 @@ theorem Complex.volume_sum_rpow_le [Nonempty ι] {p : ℝ} (hp : 1 ≤ p) (r : end LpSpace -section EuclideanSpace +namespace EuclideanSpace variable (ι : Type*) [Nonempty ι] [Fintype ι] open Fintype Real MeasureTheory MeasureTheory.Measure ENNReal -theorem EuclideanSpace.volume_ball (x : EuclideanSpace ℝ ι) (r : ℝ) : +theorem volume_ball (x : EuclideanSpace ℝ ι) (r : ℝ) : volume (Metric.ball x r) = (.ofReal r) ^ card ι * .ofReal (Real.sqrt π ^ card ι / Gamma (card ι / 2 + 1)) := by obtain hr | hr := le_total r 0 @@ -329,15 +329,15 @@ theorem EuclideanSpace.volume_ball (x : EuclideanSpace ℝ ι) (r : ℝ) : · suffices volume (Metric.ball (0 : EuclideanSpace ℝ ι) 1) = .ofReal (Real.sqrt π ^ card ι / Gamma (card ι / 2 + 1)) by rw [Measure.addHaar_ball _ _ hr, this, ofReal_pow hr, finrank_euclideanSpace] - rw [← ((EuclideanSpace.volume_preserving_measurableEquiv _).symm).measure_preimage + rw [← ((volume_preserving_measurableEquiv _).symm).measure_preimage measurableSet_ball.nullMeasurableSet] convert (volume_sum_rpow_lt_one ι one_le_two) using 4 - · simp_rw [EuclideanSpace.ball_zero_eq _ zero_le_one, one_pow, Real.rpow_two, sq_abs, + · simp_rw [ball_zero_eq _ zero_le_one, one_pow, Real.rpow_two, sq_abs, Set.setOf_app_iff] · rw [Gamma_add_one (by norm_num), Gamma_one_half_eq, ← mul_assoc, mul_div_cancel₀ _ two_ne_zero, one_mul] -theorem EuclideanSpace.volume_closedBall (x : EuclideanSpace ℝ ι) (r : ℝ) : +theorem volume_closedBall (x : EuclideanSpace ℝ ι) (r : ℝ) : volume (Metric.closedBall x r) = (.ofReal r) ^ card ι * .ofReal (sqrt π ^ card ι / Gamma (card ι / 2 + 1)) := by rw [addHaar_closedBall_eq_addHaar_ball, EuclideanSpace.volume_ball] @@ -349,14 +349,19 @@ alias Euclidean_space.volume_closedBall := EuclideanSpace.volume_closedBall end EuclideanSpace -section InnerProductSpace +namespace InnerProductSpace -open MeasureTheory MeasureTheory.Measure ENNReal Real Module +open scoped Nat +open MeasureTheory MeasureTheory.Measure ENNReal Real Module Metric variable {E : Type*} [NormedAddCommGroup E] [InnerProductSpace ℝ E] [FiniteDimensional ℝ E] - [MeasurableSpace E] [BorelSpace E] [Nontrivial E] + [MeasurableSpace E] [BorelSpace E] -theorem InnerProductSpace.volume_ball (x : E) (r : ℝ) : +section Nontrivial + +variable [Nontrivial E] + +theorem volume_ball (x : E) (r : ℝ) : volume (Metric.ball x r) = (.ofReal r) ^ finrank ℝ E * .ofReal (sqrt π ^ finrank ℝ E / Gamma (finrank ℝ E / 2 + 1)) := by rw [← ((stdOrthonormalBasis ℝ E).measurePreserving_repr_symm).measure_preimage @@ -367,13 +372,70 @@ theorem InnerProductSpace.volume_ball (x : E) (r : ℝ) : convert this simp only [LinearIsometryEquiv.preimage_ball, LinearIsometryEquiv.symm_symm, _root_.map_zero] -theorem InnerProductSpace.volume_closedBall (x : E) (r : ℝ) : +theorem volume_closedBall (x : E) (r : ℝ) : volume (Metric.closedBall x r) = (.ofReal r) ^ finrank ℝ E * .ofReal (sqrt π ^ finrank ℝ E / Gamma (finrank ℝ E / 2 + 1)) := by rw [addHaar_closedBall_eq_addHaar_ball, InnerProductSpace.volume_ball _] +lemma volume_ball_of_dim_even {k : ℕ} (hk : finrank ℝ E = 2 * k) (x : E) (r : ℝ) : + volume (ball x r) = .ofReal r ^ finrank ℝ E * .ofReal (π ^ k / (k : ℕ)!) := by + rw [volume_ball, hk, pow_mul, pow_mul, sq_sqrt pi_nonneg] + congr + simp [Gamma_nat_eq_factorial] + +lemma volume_closedBall_of_dim_even {k : ℕ} (hk : finrank ℝ E = 2 * k) (x : E) (r : ℝ) : + volume (closedBall x r) = .ofReal r ^ finrank ℝ E * .ofReal (π ^ k / (k : ℕ)!) := by + rw [addHaar_closedBall_eq_addHaar_ball, volume_ball_of_dim_even hk x] + +end Nontrivial + +lemma volume_ball_of_dim_odd {k : ℕ} (hk : finrank ℝ E = 2 * k + 1) (x : E) (r : ℝ) : + volume (ball x r) = + .ofReal r ^ finrank ℝ E * .ofReal (π ^ k * 2 ^ (k + 1) / (finrank ℝ E : ℕ)‼) := by + have : Nontrivial E := Module.nontrivial_of_finrank_pos (R := ℝ) (hk ▸ (2 * k).succ_pos) + rw [volume_ball, hk, pow_succ (√π), pow_mul, sq_sqrt pi_nonneg, mul_div_assoc, mul_div_assoc] + congr 3 + simp? [add_div, add_right_comm, -one_div, Gamma_nat_add_one_add_half] says + simp only [Nat.cast_add, Nat.cast_mul, Nat.cast_ofNat, Nat.cast_one, add_div, ne_eq, + OfNat.ofNat_ne_zero, not_false_eq_true, mul_div_cancel_left₀, add_right_comm, + Gamma_nat_add_one_add_half] + field_simp + ring + +lemma volume_closedBall_of_dim_odd {k : ℕ} (hk : finrank ℝ E = 2 * k + 1) (x : E) (r : ℝ) : + volume (closedBall x r) = + .ofReal r ^ finrank ℝ E * .ofReal (π ^ k * 2 ^ (k + 1) / (finrank ℝ E : ℕ)‼) := by + have : Nontrivial E := Module.nontrivial_of_finrank_pos (R := ℝ) (hk ▸ (2 * k).succ_pos) + rw [addHaar_closedBall_eq_addHaar_ball, volume_ball_of_dim_odd hk x r] + end InnerProductSpace +namespace EuclideanSpace + +open Real MeasureTheory MeasureTheory.Measure ENNReal Metric + +@[simp] +lemma volume_ball_fin_two (x : EuclideanSpace ℝ (Fin 2)) (r : ℝ) : + volume (ball x r) = .ofReal r ^ 2 * .ofReal π := by + norm_num [InnerProductSpace.volume_ball_of_dim_even (k := 1) (by simp) x] + +@[simp] +lemma volume_closedBall_fin_two (x : EuclideanSpace ℝ (Fin 2)) (r : ℝ) : + volume (closedBall x r) = .ofReal r ^ 2 * .ofReal π := by + rw [addHaar_closedBall_eq_addHaar_ball, volume_ball_fin_two x r] + +@[simp] +lemma volume_ball_fin_three (x : EuclideanSpace ℝ (Fin 3)) (r : ℝ) : + volume (ball x r) = .ofReal r ^ 3 * .ofReal (π * 4 / 3) := by + norm_num [InnerProductSpace.volume_ball_of_dim_odd (k := 1) (by simp) x] + +@[simp] +lemma volume_closedBall_fin_three (x : EuclideanSpace ℝ (Fin 3)) (r : ℝ) : + volume (closedBall x r) = .ofReal r ^ 3 * .ofReal (π * 4 / 3) := by + rw [addHaar_closedBall_eq_addHaar_ball, volume_ball_fin_three x] + +end EuclideanSpace + section Complex open MeasureTheory MeasureTheory.Measure ENNReal @@ -381,8 +443,7 @@ open MeasureTheory MeasureTheory.Measure ENNReal @[simp] theorem Complex.volume_ball (a : ℂ) (r : ℝ) : volume (Metric.ball a r) = .ofReal r ^ 2 * NNReal.pi := by - rw [InnerProductSpace.volume_ball a r, finrank_real_complex, Nat.cast_ofNat, div_self two_ne_zero, - one_add_one_eq_two, Real.Gamma_two, div_one, Real.sq_sqrt (by positivity), + simp [InnerProductSpace.volume_ball_of_dim_even (k := 1) (by simp) a, ← NNReal.coe_real_pi, ofReal_coe_nnreal] @[simp] diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean index 41f6d3125f98f..ccf3498aa2d2a 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean @@ -6,6 +6,7 @@ Authors: Xavier Roblot import Mathlib.MeasureTheory.Group.GeometryOfNumbers import Mathlib.MeasureTheory.Measure.Lebesgue.VolumeOfBalls import Mathlib.NumberTheory.NumberField.CanonicalEmbedding.Basic +import Mathlib.Analysis.SpecialFunctions.Gamma.BohrMollerup /-! # Convex Bodies From 1ed7634f46ba697f891ebfb3577230329d4b7196 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Sat, 14 Dec 2024 17:49:07 +0000 Subject: [PATCH 716/829] refactor(GroupTheory/SpecificGroups/Cyclic): Extract `MulEquiv` from `IsCyclic.card_mulAut` (#19948) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR extracts the isomorphism `MulAut G ≃* (ZMod (Nat.card G))ˣ` from `IsCyclic.card_mulAut`. --- Mathlib/GroupTheory/SpecificGroups/Cyclic.lean | 12 ++++++++++-- 1 file changed, 10 insertions(+), 2 deletions(-) diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index fd2f7b649dab0..ba18539a63662 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -725,12 +725,20 @@ noncomputable def mulEquivOfPrimeCardEq {p : ℕ} [Group G] [Group G'] apply mulEquivOfCyclicCardEq exact hG.trans hH.symm +variable (G) in +/-- The automorphism group of a cyclic group is isomorphic to the multiplicative group of ZMod. -/ +@[simps!] +noncomputable def IsCyclic.mulAutMulEquiv [Group G] [h : IsCyclic G] : + MulAut G ≃* (ZMod (Nat.card G))ˣ := + ((MulAut.congr (zmodCyclicMulEquiv h)).symm.trans + (MulAutMultiplicative (ZMod (Nat.card G)))).trans (ZMod.AddAutEquivUnits (Nat.card G)) + +variable (G) in theorem IsCyclic.card_mulAut [Group G] [Finite G] [h : IsCyclic G] : Nat.card (MulAut G) = Nat.totient (Nat.card G) := by have : NeZero (Nat.card G) := ⟨Nat.card_pos.ne'⟩ rw [← ZMod.card_units_eq_totient, ← Nat.card_eq_fintype_card] - exact Nat.card_congr ((MulAut.congr (zmodCyclicMulEquiv h)).toEquiv.symm.trans - (MulEquiv.toAdditive.trans (ZMod.AddAutEquivUnits (Nat.card G)).toEquiv)) + exact Nat.card_congr (mulAutMulEquiv G) end ZMod From 8fe8834dbbef8a87e09d0ba3ef2854a643b60521 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 14 Dec 2024 19:04:27 +0000 Subject: [PATCH 717/829] doc(Algebra/Category/Ring/Basic): fix copy-paste errors (#19962) --- Mathlib/Algebra/Category/Ring/Basic.lean | 14 +++++++------- 1 file changed, 7 insertions(+), 7 deletions(-) diff --git a/Mathlib/Algebra/Category/Ring/Basic.lean b/Mathlib/Algebra/Category/Ring/Basic.lean index 0933469109033..211d81e09bdde 100644 --- a/Mathlib/Algebra/Category/Ring/Basic.lean +++ b/Mathlib/Algebra/Category/Ring/Basic.lean @@ -41,7 +41,7 @@ instance : CoeSort (SemiRingCat) (Type u) := attribute [coe] SemiRingCat.carrier /-- The object in the category of R-algebras associated to a type equipped with the appropriate -typeclasses. This is the preferred way to construct a term of `AlgebraCat R`. -/ +typeclasses. This is the preferred way to construct a term of `SemiRingCat`. -/ abbrev of (R : Type u) [Semiring R] : SemiRingCat := ⟨R⟩ @@ -180,7 +180,7 @@ instance : CoeSort (RingCat) (Type u) := attribute [coe] RingCat.carrier /-- The object in the category of R-algebras associated to a type equipped with the appropriate -typeclasses. This is the preferred way to construct a term of `AlgebraCat R`. -/ +typeclasses. This is the preferred way to construct a term of `RingCat`. -/ abbrev of (R : Type u) [Ring R] : RingCat := ⟨R⟩ @@ -224,7 +224,7 @@ lemma comp_apply {R S T : RingCat} (f : R ⟶ S) (g : S ⟶ T) (r : R) : lemma hom_ext {R S : RingCat} {f g : R ⟶ S} (hf : f.hom = g.hom) : f = g := Hom.ext hf -/-- Typecheck an `AlgHom` as a morphism in `AlgebraCat R`. -/ +/-- Typecheck a `RingHom` as a morphism in `RingCat`. -/ abbrev ofHom {R S : Type u} [Ring R] [Ring S] (f : R →+* S) : of R ⟶ of S := ⟨f⟩ @@ -319,7 +319,7 @@ instance : CoeSort (CommSemiRingCat) (Type u) := attribute [coe] CommSemiRingCat.carrier /-- The object in the category of R-algebras associated to a type equipped with the appropriate -typeclasses. This is the preferred way to construct a term of `AlgebraCat R`. -/ +typeclasses. This is the preferred way to construct a term of `CommSemiRingCat`. -/ abbrev of (R : Type u) [CommSemiring R] : CommSemiRingCat := ⟨R⟩ @@ -363,7 +363,7 @@ lemma comp_apply {R S T : CommSemiRingCat} (f : R ⟶ S) (g : S ⟶ T) (r : R) : lemma hom_ext {R S : CommSemiRingCat} {f g : R ⟶ S} (hf : f.hom = g.hom) : f = g := Hom.ext hf -/-- Typecheck an `AlgHom` as a morphism in `AlgebraCat R`. -/ +/-- Typecheck a `RingHom` as a morphism in `CommSemiRingCat`. -/ abbrev ofHom {R S : Type u} [CommSemiring R] [CommSemiring S] (f : R →+* S) : of R ⟶ of S := ⟨f⟩ @@ -461,7 +461,7 @@ instance : CoeSort (CommRingCat) (Type u) := attribute [coe] CommRingCat.carrier /-- The object in the category of R-algebras associated to a type equipped with the appropriate -typeclasses. This is the preferred way to construct a term of `AlgebraCat R`. -/ +typeclasses. This is the preferred way to construct a term of `CommRingCat`. -/ abbrev of (R : Type u) [CommRing R] : CommRingCat := ⟨R⟩ @@ -505,7 +505,7 @@ lemma comp_apply {R S T : CommRingCat} (f : R ⟶ S) (g : S ⟶ T) (r : R) : lemma hom_ext {R S : CommRingCat} {f g : R ⟶ S} (hf : f.hom = g.hom) : f = g := Hom.ext hf -/-- Typecheck an `AlgHom` as a morphism in `AlgebraCat R`. -/ +/-- Typecheck a `RingHom` as a morphism in `CommRingCat`. -/ abbrev ofHom {R S : Type u} [CommRing R] [CommRing S] (f : R →+* S) : of R ⟶ of S := ⟨f⟩ From 37a149f02b44682bb186c430d0a896dce3c1f17f Mon Sep 17 00:00:00 2001 From: smorel394 Date: Sat, 14 Dec 2024 20:37:47 +0000 Subject: [PATCH 718/829] feat(Algebra/Category/Grp/ForgetCorepresentable, Algebra/Category/MonCat/ForgetCorepresentable): the forgetful functor on the category of monoids is corepresentable (#19964) The forgetful functor on the various categories of monoids (additive/multiplicative/commutative/etc) are representable by the natural numbers. This basically copies the corresponding code for the group categories, and moves two lemmas from `Algebra.Category.Grp.ForgetCorepresentable` to `Algebra.Category.MonCat.ForgetCorepresentable`, to avoid having a `MonCat` file importing a `Grp` file. Will be used to proved that to prove that the forgetful functor on `MonCat` creates all limits. --- Mathlib.lean | 1 + .../Category/Grp/ForgetCorepresentable.lean | 21 +--- .../MonCat/ForgetCorepresentable.lean | 112 ++++++++++++++++++ 3 files changed, 114 insertions(+), 20 deletions(-) create mode 100644 Mathlib/Algebra/Category/MonCat/ForgetCorepresentable.lean diff --git a/Mathlib.lean b/Mathlib.lean index f26acaca225ac..87d14e789d6f0 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -143,6 +143,7 @@ import Mathlib.Algebra.Category.MonCat.Adjunctions import Mathlib.Algebra.Category.MonCat.Basic import Mathlib.Algebra.Category.MonCat.Colimits import Mathlib.Algebra.Category.MonCat.FilteredColimits +import Mathlib.Algebra.Category.MonCat.ForgetCorepresentable import Mathlib.Algebra.Category.MonCat.Limits import Mathlib.Algebra.Category.Ring.Adjunctions import Mathlib.Algebra.Category.Ring.Basic diff --git a/Mathlib/Algebra/Category/Grp/ForgetCorepresentable.lean b/Mathlib/Algebra/Category/Grp/ForgetCorepresentable.lean index c7caa5ef7b7ad..bf80cb3f9cbf7 100644 --- a/Mathlib/Algebra/Category/Grp/ForgetCorepresentable.lean +++ b/Mathlib/Algebra/Category/Grp/ForgetCorepresentable.lean @@ -6,6 +6,7 @@ Authors: Joël Riou import Mathlib.Algebra.Category.Grp.Basic import Mathlib.Algebra.Group.ULift import Mathlib.CategoryTheory.Yoneda +import Mathlib.Algebra.Category.MonCat.ForgetCorepresentable /-! # The forget functor is corepresentable @@ -22,16 +23,6 @@ open CategoryTheory Opposite namespace MonoidHom -/-- The equivalence `(β →* γ) ≃ (α →* γ)` obtained by precomposition with -a multiplicative equivalence `e : α ≃* β`. -/ -@[simps] -def precompEquiv {α β : Type*} [Monoid α] [Monoid β] (e : α ≃* β) (γ : Type*) [Monoid γ] : - (β →* γ) ≃ (α →* γ) where - toFun f := f.comp e - invFun g := g.comp e.symm - left_inv _ := by ext; simp - right_inv _ := by ext; simp - /-- The equivalence `(Multiplicative ℤ →* α) ≃ α` for any group `α`. -/ @[simps] def fromMultiplicativeIntEquiv (α : Type u) [Group α] : (Multiplicative ℤ →* α) ≃ α where @@ -50,16 +41,6 @@ end MonoidHom namespace AddMonoidHom -/-- The equivalence `(β →+ γ) ≃ (α →+ γ)` obtained by precomposition with -an additive equivalence `e : α ≃+ β`. -/ -@[simps] -def precompEquiv {α β : Type*} [AddMonoid α] [AddMonoid β] (e : α ≃+ β) (γ : Type*) [AddMonoid γ] : - (β →+ γ) ≃ (α →+ γ) where - toFun f := f.comp e - invFun g := g.comp e.symm - left_inv _ := by ext; simp - right_inv _ := by ext; simp - /-- The equivalence `(ℤ →+ α) ≃ α` for any additive group `α`. -/ @[simps] def fromIntEquiv (α : Type u) [AddGroup α] : (ℤ →+ α) ≃ α where diff --git a/Mathlib/Algebra/Category/MonCat/ForgetCorepresentable.lean b/Mathlib/Algebra/Category/MonCat/ForgetCorepresentable.lean new file mode 100644 index 0000000000000..796d9ab458bb6 --- /dev/null +++ b/Mathlib/Algebra/Category/MonCat/ForgetCorepresentable.lean @@ -0,0 +1,112 @@ +/- +Copyright (c) 2024 Sophie Morel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sophie Morel +-/ +import Mathlib.Algebra.Category.MonCat.Basic +import Mathlib.Data.Nat.Cast.Basic + +/-! +# The forget functor is corepresentable + +The forgetful functor `AddCommMonCat.{u} ⥤ Type u` is corepresentable +by `ULift ℕ`. Similar results are obtained for the variants `CommMonCat`, `AddMonCat` +and `MonCat`. + +-/ + +universe u + +open CategoryTheory Opposite + +namespace MonoidHom + +/-- The equivalence `(β →* γ) ≃ (α →* γ)` obtained by precomposition with +a multiplicative equivalence `e : α ≃* β`. -/ +@[simps] +def precompEquiv {α β : Type*} [Monoid α] [Monoid β] (e : α ≃* β) (γ : Type*) [Monoid γ] : + (β →* γ) ≃ (α →* γ) where + toFun f := f.comp e + invFun g := g.comp e.symm + left_inv _ := by ext; simp + right_inv _ := by ext; simp + +/-- The equivalence `(Multiplicative ℕ →* α) ≃ α` for any monoid `α`. -/ +@[simps] +def fromMultiplicativeNatEquiv (α : Type u) [Monoid α] : (Multiplicative ℕ →* α) ≃ α where + toFun φ := φ (Multiplicative.ofAdd 1) + invFun x := powersHom α x + left_inv φ := by ext; simp + right_inv x := by simp + +/-- The equivalence `(ULift (Multiplicative ℕ) →* α) ≃ α` for any monoid `α`. -/ +@[simps!] +def fromULiftMultiplicativeNatEquiv (α : Type u) [Monoid α] : + (ULift.{u} (Multiplicative ℕ) →* α) ≃ α := + (precompEquiv (MulEquiv.ulift.symm) _).trans (fromMultiplicativeNatEquiv α) + +end MonoidHom + +namespace AddMonoidHom + +/-- The equivalence `(β →+ γ) ≃ (α →+ γ)` obtained by precomposition with +an additive equivalence `e : α ≃+ β`. -/ +@[simps] +def precompEquiv {α β : Type*} [AddMonoid α] [AddMonoid β] (e : α ≃+ β) (γ : Type*) [AddMonoid γ] : + (β →+ γ) ≃ (α →+ γ) where + toFun f := f.comp e + invFun g := g.comp e.symm + left_inv _ := by ext; simp + right_inv _ := by ext; simp + +/-- The equivalence `(ℤ →+ α) ≃ α` for any additive group `α`. -/ +@[simps] +def fromNatEquiv (α : Type u) [AddMonoid α] : (ℕ →+ α) ≃ α where + toFun φ := φ 1 + invFun x := multiplesHom α x + left_inv φ := by ext; simp + right_inv x := by simp + +/-- The equivalence `(ULift ℕ →+ α) ≃ α` for any additive monoid `α`. -/ +@[simps!] +def fromULiftNatEquiv (α : Type u) [AddMonoid α] : (ULift.{u} ℕ →+ α) ≃ α := + (precompEquiv (AddEquiv.ulift.symm) _).trans (fromNatEquiv α) + +end AddMonoidHom + +/-- The forgetful functor `MonCat.{u} ⥤ Type u` is corepresentable. -/ +def MonCat.coyonedaObjIsoForget : + coyoneda.obj (op (of (ULift.{u} (Multiplicative ℕ)))) ≅ forget MonCat.{u} := + (NatIso.ofComponents (fun M => (MonoidHom.fromULiftMultiplicativeNatEquiv M.α).toIso)) + + +/-- The forgetful functor `CommMonCat.{u} ⥤ Type u` is corepresentable. -/ +def CommMonCat.coyonedaObjIsoForget : + coyoneda.obj (op (of (ULift.{u} (Multiplicative ℕ)))) ≅ forget CommMonCat.{u} := + (NatIso.ofComponents (fun M => (MonoidHom.fromULiftMultiplicativeNatEquiv M.α).toIso)) + +/-- The forgetful functor `AddMonCat.{u} ⥤ Type u` is corepresentable. -/ +def AddMonCat.coyonedaObjIsoForget : + coyoneda.obj (op (of (ULift.{u} ℕ))) ≅ forget AddMonCat.{u} := + (NatIso.ofComponents (fun M => (AddMonoidHom.fromULiftNatEquiv M.α).toIso)) + +/-- The forgetful functor `AddCommMonCat.{u} ⥤ Type u` is corepresentable. -/ +def AddCommMonCat.coyonedaObjIsoForget : + coyoneda.obj (op (of (ULift.{u} ℕ))) ≅ forget AddCommMonCat.{u} := + (NatIso.ofComponents (fun M => (AddMonoidHom.fromULiftNatEquiv M.α).toIso)) + +instance MonCat.forget_isCorepresentable : + (forget MonCat.{u}).IsCorepresentable := + Functor.IsCorepresentable.mk' MonCat.coyonedaObjIsoForget + +instance CommMonCat.forget_isCorepresentable : + (forget CommMonCat.{u}).IsCorepresentable := + Functor.IsCorepresentable.mk' CommMonCat.coyonedaObjIsoForget + +instance AddMonCat.forget_isCorepresentable : + (forget AddMonCat.{u}).IsCorepresentable := + Functor.IsCorepresentable.mk' AddMonCat.coyonedaObjIsoForget + +instance AddCommMonCat.forget_isCorepresentable : + (forget AddCommMonCat.{u}).IsCorepresentable := + Functor.IsCorepresentable.mk' AddCommMonCat.coyonedaObjIsoForget From ecf9698f24796d565639fd0cfef57a8ac88876ff Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sun, 15 Dec 2024 00:03:12 +0000 Subject: [PATCH 719/829] chore(SetTheory/ZFC/Rank): clean up file (#18239) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR does the following: - Make various theorems take explicit arguments (as it's a bit painful to invoke them otherwise) - Replace `lsub f` by `⨆ i, succ (f i)`, in anticipation for the former being deprecated. --- Mathlib/SetTheory/ZFC/Rank.lean | 79 ++++++++++++++++++--------------- 1 file changed, 42 insertions(+), 37 deletions(-) diff --git a/Mathlib/SetTheory/ZFC/Rank.lean b/Mathlib/SetTheory/ZFC/Rank.lean index 944b56ee03b14..0aaa4d9afbda4 100644 --- a/Mathlib/SetTheory/ZFC/Rank.lean +++ b/Mathlib/SetTheory/ZFC/Rank.lean @@ -23,33 +23,36 @@ universe u v open Ordinal Order +/-! ### PSet rank -/ + namespace PSet /-- The ordinal rank of a pre-set -/ noncomputable def rank : PSet.{u} → Ordinal.{u} - | ⟨_, A⟩ => lsub fun a => rank (A a) + | ⟨_, A⟩ => ⨆ a, succ (rank (A a)) theorem rank_congr : ∀ {x y : PSet}, Equiv x y → rank x = rank y - | ⟨_, _⟩, ⟨_, _⟩, ⟨αβ, βα⟩ => - lsub_eq_of_range_eq (by - ext - constructor <;> simp <;> intro a h - · obtain ⟨b, h'⟩ := αβ a - exists b - rw [← h, rank_congr h'] - · obtain ⟨b, h'⟩ := βα a - exists b - rw [← h, rank_congr h']) + | ⟨_, _⟩, ⟨_, _⟩, ⟨αβ, βα⟩ => by + apply congr_arg sSup + ext + constructor <;> simp <;> intro a h + · obtain ⟨b, h'⟩ := αβ a + exists b + rw [← h, rank_congr h'] + · obtain ⟨b, h'⟩ := βα a + exists b + rw [← h, rank_congr h'] theorem rank_lt_of_mem : ∀ {x y : PSet}, y ∈ x → rank y < rank x | ⟨_, _⟩, _, ⟨_, h⟩ => by - rw [rank_congr h] - apply lt_lsub + rw [rank_congr h, ← succ_le_iff] + apply Ordinal.le_iSup theorem rank_le_iff {o : Ordinal} : ∀ {x : PSet}, rank x ≤ o ↔ ∀ ⦃y⦄, y ∈ x → rank y < o - | ⟨_, A⟩ => - ⟨fun h _ h' => (rank_lt_of_mem h').trans_le h, fun h => - lsub_le fun a => h (Mem.mk A a)⟩ + | ⟨_, A⟩ => by + refine ⟨fun h _ h' => (rank_lt_of_mem h').trans_le h, fun h ↦ Ordinal.iSup_le fun a ↦ ?_⟩ + rw [succ_le_iff] + exact h (Mem.mk A a) theorem lt_rank_iff {o : Ordinal} {x : PSet} : o < rank x ↔ ∃ y ∈ x, o ≤ rank y := by rw [← not_iff_not, not_lt, rank_le_iff] @@ -64,7 +67,7 @@ variable {x y : PSet.{u}} theorem rank_empty : rank ∅ = 0 := by simp [rank] @[simp] -theorem rank_insert : rank (insert x y) = max (succ (rank x)) (rank y) := by +theorem rank_insert (x y : PSet) : rank (insert x y) = max (succ (rank x)) (rank y) := by apply le_antisymm · simp_rw [rank_le_iff, mem_insert_iff] rintro _ (h | h) @@ -75,14 +78,14 @@ theorem rank_insert : rank (insert x y) = max (succ (rank x)) (rank y) := by · exact rank_mono (subset_iff.2 fun z => mem_insert_of_mem x) @[simp] -theorem rank_singleton : rank {x} = succ (rank x) := - rank_insert.trans (by simp) +theorem rank_singleton (x : PSet) : rank {x} = succ (rank x) := + (rank_insert _ _).trans (by simp) -theorem rank_pair : rank {x, y} = max (succ (rank x)) (succ (rank y)) := by +theorem rank_pair (x y : PSet) : rank {x, y} = max (succ (rank x)) (succ (rank y)) := by simp @[simp] -theorem rank_powerset : rank (powerset x) = succ (rank x) := by +theorem rank_powerset (x : PSet) : rank (powerset x) = succ (rank x) := by apply le_antisymm · simp_rw [rank_le_iff, mem_powerset, lt_succ_iff] intro @@ -94,12 +97,12 @@ theorem rank_powerset : rank (powerset x) = succ (rank x) := by /-- For the rank of `⋃₀ x`, we only have `rank (⋃₀ x) ≤ rank x ≤ rank (⋃₀ x) + 1`. This inequality is split into `rank_sUnion_le` and `le_succ_rank_sUnion`. -/ -theorem rank_sUnion_le : rank (⋃₀ x) ≤ rank x := by +theorem rank_sUnion_le (x : PSet) : rank (⋃₀ x) ≤ rank x := by simp_rw [rank_le_iff, mem_sUnion] intro _ ⟨_, _, _⟩ trans <;> apply rank_lt_of_mem <;> assumption -theorem le_succ_rank_sUnion : rank x ≤ succ (rank (⋃₀ x)) := by +theorem le_succ_rank_sUnion (x : PSet) : rank x ≤ succ (rank (⋃₀ x)) := by rw [← rank_powerset] apply rank_mono rw [subset_iff] @@ -122,6 +125,8 @@ theorem rank_eq_wfRank : lift.{u + 1, u} (rank x) = IsWellFounded.rank (α := PS end PSet +/-! ### ZFSet rank -/ + namespace ZFSet variable {x y : ZFSet.{u}} @@ -149,18 +154,18 @@ theorem lt_rank_iff {o : Ordinal} : o < rank x ↔ ∃ y ∈ x, o ≤ rank y := theorem rank_empty : rank ∅ = 0 := PSet.rank_empty @[simp] -theorem rank_insert : rank (insert x y) = max (succ (rank x)) (rank y) := - Quotient.inductionOn₂ x y fun _ _ => PSet.rank_insert +theorem rank_insert (x y : ZFSet) : rank (insert x y) = max (succ (rank x)) (rank y) := + Quotient.inductionOn₂ x y PSet.rank_insert @[simp] -theorem rank_singleton : rank {x} = succ (rank x) := - rank_insert.trans (by simp) +theorem rank_singleton (x : ZFSet) : rank {x} = succ (rank x) := + (rank_insert _ _).trans (by simp) -theorem rank_pair : rank {x, y} = max (succ (rank x)) (succ (rank y)) := by +theorem rank_pair (x y : ZFSet) : rank {x, y} = max (succ (rank x)) (succ (rank y)) := by simp @[simp] -theorem rank_union : rank (x ∪ y) = max (rank x) (rank y) := by +theorem rank_union (x y : ZFSet) : rank (x ∪ y) = max (rank x) (rank y) := by apply le_antisymm · simp_rw [rank_le_iff, mem_union, lt_max_iff] intro @@ -168,18 +173,18 @@ theorem rank_union : rank (x ∪ y) = max (rank x) (rank y) := by · apply max_le <;> apply rank_mono <;> intro _ h <;> simp [h] @[simp] -theorem rank_powerset : rank (powerset x) = succ (rank x) := - Quotient.inductionOn x fun _ => PSet.rank_powerset +theorem rank_powerset (x : ZFSet) : rank (powerset x) = succ (rank x) := + Quotient.inductionOn x PSet.rank_powerset /-- For the rank of `⋃₀ x`, we only have `rank (⋃₀ x) ≤ rank x ≤ rank (⋃₀ x) + 1`. This inequality is split into `rank_sUnion_le` and `le_succ_rank_sUnion`. -/ -theorem rank_sUnion_le : rank (⋃₀ x) ≤ rank x := by +theorem rank_sUnion_le (x : ZFSet) : rank (⋃₀ x) ≤ rank x := by simp_rw [rank_le_iff, mem_sUnion] intro _ ⟨_, _, _⟩ trans <;> apply rank_lt_of_mem <;> assumption -theorem le_succ_rank_sUnion : rank x ≤ succ (rank (⋃₀ x)) := by +theorem le_succ_rank_sUnion (x : ZFSet) : rank x ≤ succ (rank (⋃₀ x)) := by rw [← rank_powerset] apply rank_mono intro z _ @@ -189,10 +194,10 @@ theorem le_succ_rank_sUnion : rank x ≤ succ (rank (⋃₀ x)) := by exists z @[simp] -theorem rank_range {α : Type u} {f : α → ZFSet.{max u v}} : - rank (range f) = lsub fun i => rank (f i) := by - apply (lsub_le _).antisymm' - · simpa [rank_le_iff] using lt_lsub _ +theorem rank_range {α : Type u} (f : α → ZFSet.{max u v}) : + rank (range f) = ⨆ i, succ (rank (f i)) := by + apply (Ordinal.iSup_le _).antisymm' + · simpa [rank_le_iff, ← succ_le_iff] using Ordinal.le_iSup _ · simp [rank_lt_of_mem] /-- `ZFSet.rank` is equal to the `IsWellFounded.rank` over `∈`. -/ From 86f77dba34177ba6a7b9a7843d943abbafeb5ab3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sun, 15 Dec 2024 00:11:31 +0000 Subject: [PATCH 720/829] =?UTF-8?q?feat(SetTheory/ZFC/Basic):=20`x=20?= =?UTF-8?q?=E2=88=88=20y=20=E2=86=92=20=C2=AC=20y=20=E2=8A=86=20x`=20(#199?= =?UTF-8?q?67)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/SetTheory/ZFC/Basic.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/Mathlib/SetTheory/ZFC/Basic.lean b/Mathlib/SetTheory/ZFC/Basic.lean index 535579e53318c..7f81281646749 100644 --- a/Mathlib/SetTheory/ZFC/Basic.lean +++ b/Mathlib/SetTheory/ZFC/Basic.lean @@ -254,6 +254,12 @@ theorem mem_asymm {x y : PSet} : x ∈ y → y ∉ x := theorem mem_irrefl (x : PSet) : x ∉ x := irrefl_of (· ∈ ·) x +theorem not_subset_of_mem {x y : PSet} (h : x ∈ y) : ¬ y ⊆ x := + fun h' ↦ mem_irrefl _ <| mem_of_subset h' h + +theorem not_mem_of_subset {x y : PSet} (h : x ⊆ y) : y ∉ x := + imp_not_comm.2 not_subset_of_mem h + /-- Convert a pre-set to a `Set` of pre-sets. -/ def toSet (u : PSet.{u}) : Set PSet.{u} := { x | x ∈ u } @@ -1159,6 +1165,12 @@ theorem mem_asymm {x y : ZFSet} : x ∈ y → y ∉ x := theorem mem_irrefl (x : ZFSet) : x ∉ x := irrefl_of (· ∈ ·) x +theorem not_subset_of_mem {x y : ZFSet} (h : x ∈ y) : ¬ y ⊆ x := + fun h' ↦ mem_irrefl _ (h' h) + +theorem not_mem_of_subset {x y : ZFSet} (h : x ⊆ y) : y ∉ x := + imp_not_comm.2 not_subset_of_mem h + theorem regularity (x : ZFSet.{u}) (h : x ≠ ∅) : ∃ y ∈ x, x ∩ y = ∅ := by_contradiction fun ne => h <| (eq_empty x).2 fun y => From a395cd00a4fbe946ca76ff1fd1d1c78129a135ff Mon Sep 17 00:00:00 2001 From: leanprover-community-bot-assistant Date: Sun, 15 Dec 2024 00:20:11 +0000 Subject: [PATCH 721/829] chore(scripts): update nolints.json (#19971) I am happy to remove some nolints for you! --- scripts/nolints.json | 2 -- 1 file changed, 2 deletions(-) diff --git a/scripts/nolints.json b/scripts/nolints.json index 58730127f796d..6e98e387e7e8d 100644 --- a/scripts/nolints.json +++ b/scripts/nolints.json @@ -281,7 +281,6 @@ ["docBlame", "ProbabilityTheory.«termEVar[_]»"], ["docBlame", "ProbabilityTheory.«termVar[_]»"], ["docBlame", "ProbabilityTheory.«term_=ₐₛ_»"], - ["docBlame", "ProbabilityTheory.«term_[_]»"], ["docBlame", "ProbabilityTheory.«term_×ₖ_»"], ["docBlame", "ProbabilityTheory.«term_≤ₐₛ_»"], ["docBlame", "ProbabilityTheory.«term_⟦_|_⟧»"], @@ -295,7 +294,6 @@ ["docBlame", "QPF.abs"], ["docBlame", "QPF.repr"], ["docBlame", "QuadraticMap.toFun"], - ["docBlame", "Quaternion.«termℍ[_]»"], ["docBlame", "RCLike.im"], ["docBlame", "RCLike.re"], ["docBlame", "Random.randBool"], From 161b962fd5670f7d764c498611b9870c1405f4c6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sun, 15 Dec 2024 00:44:26 +0000 Subject: [PATCH 722/829] feat(SetTheory/ZFC/Ordinal): Alternate characterizations of ordinals (#17001) We previously defined ordinals as transitive sets, transitive under membership. This PR establishes this definition is equivalent to the following: - An ordinal is a transitive set of transitive sets - An ordinal is a transitive set of ordinals - An ordinal is a transitive set, trichotomous under membership - An ordinal is a transitive set, well-ordered under membership --- Mathlib/Order/RelClasses.lean | 9 ++ Mathlib/Order/RelIso/Set.lean | 15 +++ Mathlib/SetTheory/ZFC/Ordinal.lean | 152 +++++++++++++++++++++++++---- 3 files changed, 157 insertions(+), 19 deletions(-) diff --git a/Mathlib/Order/RelClasses.lean b/Mathlib/Order/RelClasses.lean index 0acae0e49466b..9fe8af7f08c81 100644 --- a/Mathlib/Order/RelClasses.lean +++ b/Mathlib/Order/RelClasses.lean @@ -234,6 +234,11 @@ theorem WellFoundedRelation.asymmetric {α : Sort*} [WellFoundedRelation α] {a fun hab hba => WellFoundedRelation.asymmetric hba hab termination_by a +theorem WellFoundedRelation.asymmetric₃ {α : Sort*} [WellFoundedRelation α] {a b c : α} : + WellFoundedRelation.rel a b → WellFoundedRelation.rel b c → ¬ WellFoundedRelation.rel c a := + fun hab hbc hca => WellFoundedRelation.asymmetric₃ hca hab hbc +termination_by a + lemma WellFounded.prod_lex {ra : α → α → Prop} {rb : β → β → Prop} (ha : WellFounded ra) (hb : WellFounded rb) : WellFounded (Prod.Lex ra rb) := (Prod.lex ⟨_, ha⟩ ⟨_, hb⟩).wf @@ -295,6 +300,10 @@ theorem WellFounded.asymmetric {α : Sort*} {r : α → α → Prop} (h : WellFo r a b → ¬r b a := @WellFoundedRelation.asymmetric _ ⟨_, h⟩ _ _ +theorem WellFounded.asymmetric₃ {α : Sort*} {r : α → α → Prop} (h : WellFounded r) (a b c) : + r a b → r b c → ¬r c a := + @WellFoundedRelation.asymmetric₃ _ ⟨_, h⟩ _ _ _ + -- see Note [lower instance priority] instance (priority := 100) (r : α → α → Prop) [IsWellFounded α r] : IsAsymm α r := ⟨IsWellFounded.wf.asymmetric⟩ diff --git a/Mathlib/Order/RelIso/Set.lean b/Mathlib/Order/RelIso/Set.lean index 583d31312e94a..e7e04319e0243 100644 --- a/Mathlib/Order/RelIso/Set.lean +++ b/Mathlib/Order/RelIso/Set.lean @@ -63,6 +63,18 @@ protected def relEmbedding (r : α → α → Prop) (p : Set α) : Subrel r p theorem relEmbedding_apply (r : α → α → Prop) (p a) : Subrel.relEmbedding r p a = a.1 := rfl +/-- A set inclusion as a relation embedding. -/ +protected def inclusionEmbedding (r : α → α → Prop) {p q : Set α} (h : p ⊆ q) : + Subrel r p ↪r Subrel r q where + toFun := Set.inclusion h + inj' _ _ h := (Set.inclusion_inj _).mp h + map_rel_iff' := Iff.rfl + +@[simp] +theorem coe_inclusionEmbedding (r : α → α → Prop) {p q : Set α} (h : p ⊆ q) : + (Subrel.inclusionEmbedding r h : p → q) = Set.inclusion h := + rfl + instance (r : α → α → Prop) [IsWellOrder α r] (p : Set α) : IsWellOrder p (Subrel r p) := RelEmbedding.isWellOrder (Subrel.relEmbedding r p) @@ -80,6 +92,9 @@ instance (r : α → α → Prop) [IsRefl α r] (p : Set α) : IsRefl p (Subrel instance (r : α → α → Prop) [IsSymm α r] (p : Set α) : IsSymm p (Subrel r p) := ⟨fun x y => @IsSymm.symm α r _ x y⟩ +instance (r : α → α → Prop) [IsAsymm α r] (p : Set α) : IsAsymm p (Subrel r p) := + ⟨fun x y => @IsAsymm.asymm α r _ x y⟩ + instance (r : α → α → Prop) [IsTrans α r] (p : Set α) : IsTrans p (Subrel r p) := ⟨fun x y z => @IsTrans.trans α r _ x y z⟩ diff --git a/Mathlib/SetTheory/ZFC/Ordinal.lean b/Mathlib/SetTheory/ZFC/Ordinal.lean index 43a8b4f282040..ae1345400e7b7 100644 --- a/Mathlib/SetTheory/ZFC/Ordinal.lean +++ b/Mathlib/SetTheory/ZFC/Ordinal.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Violeta Hernández Palacios. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Violeta Hernández Palacios -/ +import Mathlib.Order.GameAdd import Mathlib.Order.RelIso.Set import Mathlib.SetTheory.ZFC.Basic @@ -12,27 +13,25 @@ import Mathlib.SetTheory.ZFC.Basic This file works towards the development of von Neumann ordinals, i.e. transitive sets, well-ordered under `∈`. -We currently only have an initial development of transitive sets and ordinals. Further development -can be found on the Mathlib 3 branch `von_neumann_v2`. - ## Definitions - `ZFSet.IsTransitive` means that every element of a set is a subset. -- `ZFSet.IsOrdinal` means that the set is transitive and well-ordered under `∈`. +- `ZFSet.IsOrdinal` means that the set is transitive and well-ordered under `∈`. We show multiple + equivalences to this definition. ## TODO -- Define von Neumann ordinals and the von Neumann hierarchy. -- Build correspondences between these notions and those of the standard `Ordinal` type. +- Define the von Neumann hierarchy. +- Build correspondences between these set notions and those of the standard `Ordinal` type. -/ - universe u -variable {x y z : ZFSet.{u}} +variable {x y z w : ZFSet.{u}} namespace ZFSet +/-! ### Transitive sets -/ /-- A transitive set is one where every element is a subset. @@ -98,8 +97,10 @@ theorem isTransitive_iff_subset_powerset : x.IsTransitive ↔ x ⊆ powerset x : alias ⟨IsTransitive.subset_powerset, _⟩ := isTransitive_iff_subset_powerset +/-! ### Ordinals -/ + /-- A set `x` is a von Neumann ordinal when it's a transitive set, that's transitive under `∈`. We -will prove that this further implies that `x` is well-ordered under `∈`. +prove that this further implies that `x` is well-ordered under `∈` in `isOrdinal_iff_isWellOrder`. The transitivity condition `a ∈ b → b ∈ c → a ∈ c` can be written without assuming `a ∈ x` and `b ∈ x`. The lemma `isOrdinal_iff_isTrans` shows this condition is equivalent to the usual one. -/ @@ -120,17 +121,130 @@ theorem mem_trans (h : z.IsOrdinal) : x ∈ y → y ∈ z → x ∈ z := protected theorem isTrans (h : x.IsOrdinal) : IsTrans x.toSet (Subrel (· ∈ ·) _) := ⟨fun _ _ c hab hbc => h.mem_trans' hab hbc c.2⟩ -end IsOrdinal - /-- The simplified form of transitivity used within `IsOrdinal` yields an equivalent definition to the standard one. -/ -theorem isOrdinal_iff_isTrans : - x.IsOrdinal ↔ x.IsTransitive ∧ IsTrans x.toSet (Subrel (· ∈ ·) _) := by - use fun h => ⟨h.isTransitive, h.isTrans⟩ - rintro ⟨h₁, ⟨h₂⟩⟩ - use h₁ - intro y z w hyz hzw hwx - have hzx := h₁.mem_trans hzw hwx - exact h₂ ⟨y, h₁.mem_trans hyz hzx⟩ ⟨z, hzx⟩ ⟨w, hwx⟩ hyz hzw +theorem _root_.ZFSet.isOrdinal_iff_isTrans : + x.IsOrdinal ↔ x.IsTransitive ∧ IsTrans x.toSet (Subrel (· ∈ ·) _) where + mp h := ⟨h.isTransitive, h.isTrans⟩ + mpr := by + rintro ⟨h₁, ⟨h₂⟩⟩ + refine ⟨h₁, @fun y z w hyz hzw hwx ↦ ?_⟩ + have hzx := h₁.mem_trans hzw hwx + exact h₂ ⟨y, h₁.mem_trans hyz hzx⟩ ⟨z, hzx⟩ ⟨w, hwx⟩ hyz hzw + +protected theorem mem (hx : x.IsOrdinal) (hy : y ∈ x) : y.IsOrdinal := + have := hx.isTrans + let f : Subrel (· ∈ ·) y.toSet ↪r Subrel (· ∈ ·) x.toSet := + Subrel.inclusionEmbedding (· ∈ ·) (hx.subset_of_mem hy) + isOrdinal_iff_isTrans.2 ⟨fun _ hz _ ha ↦ hx.mem_trans' ha hz hy, f.isTrans⟩ + +/-- An ordinal is a transitive set of transitive sets. -/ +theorem _root_.ZFSet.isOrdinal_iff_forall_mem_isTransitive : + x.IsOrdinal ↔ x.IsTransitive ∧ ∀ y ∈ x, y.IsTransitive where + mp h := ⟨h.isTransitive, fun _ hy ↦ (h.mem hy).isTransitive⟩ + mpr := fun ⟨h₁, h₂⟩ ↦ ⟨h₁, fun hyz hzw hwx ↦ (h₂ _ hwx).mem_trans hyz hzw⟩ + +/-- An ordinal is a transitive set of ordinals. -/ +theorem _root_.ZFSet.isOrdinal_iff_forall_mem_isOrdinal : + x.IsOrdinal ↔ x.IsTransitive ∧ ∀ y ∈ x, y.IsOrdinal where + mp h := ⟨h.isTransitive, fun _ ↦ h.mem⟩ + mpr := fun ⟨h₁, h₂⟩ ↦ isOrdinal_iff_forall_mem_isTransitive.2 + ⟨h₁, fun y hy ↦ (h₂ y hy).isTransitive⟩ + +theorem subset_iff_eq_or_mem (hx : x.IsOrdinal) (hy : y.IsOrdinal) : x ⊆ y ↔ x = y ∨ x ∈ y := by + constructor + · revert hx hy + apply Sym2.GameAdd.induction mem_wf _ x y + intro x y IH hx hy hxy + by_cases hyx : y ⊆ x + · exact Or.inl (subset_antisymm hxy hyx) + · obtain ⟨m, hm, hm'⟩ := mem_wf.has_min (y.toSet \ x.toSet) (Set.diff_nonempty.2 hyx) + have hmy : m ∈ y := show m ∈ y.toSet from Set.mem_of_mem_diff hm + have hmx : m ⊆ x := by + intro z hzm + by_contra hzx + exact hm' _ ⟨hy.mem_trans hzm hmy, hzx⟩ hzm + obtain rfl | H := IH m x (Sym2.GameAdd.fst_snd hmy) (hy.mem hmy) hx hmx + · exact Or.inr hmy + · cases Set.not_mem_of_mem_diff hm H + · rintro (rfl | h) + · rfl + · exact hy.subset_of_mem h + +alias ⟨eq_or_mem_of_subset, _⟩ := subset_iff_eq_or_mem + +theorem mem_of_subset_of_mem (h : x.IsOrdinal) (hz : z.IsOrdinal) (hx : x ⊆ y) (hy : y ∈ z) : + x ∈ z := by + obtain rfl | hx := h.eq_or_mem_of_subset (hz.mem hy) hx + · exact hy + · exact hz.mem_trans hx hy + +theorem not_mem_iff_subset (hx : x.IsOrdinal) (hy : y.IsOrdinal) : x ∉ y ↔ y ⊆ x := by + refine ⟨?_, fun hxy hyx ↦ mem_irrefl _ (hxy hyx)⟩ + revert hx hy + apply Sym2.GameAdd.induction mem_wf _ x y + intros x y IH hx hy hyx z hzy + by_contra hzx + exact hyx (mem_of_subset_of_mem hx hy (IH z x (Sym2.GameAdd.fst_snd hzy) (hy.mem hzy) hx hzx) hzy) + +theorem not_subset_iff_mem (hx : x.IsOrdinal) (hy : y.IsOrdinal) : ¬ x ⊆ y ↔ y ∈ x := by + rw [not_iff_comm, not_mem_iff_subset hy hx] + +theorem mem_or_subset (hx : x.IsOrdinal) (hy : y.IsOrdinal) : x ∈ y ∨ y ⊆ x := by + rw [or_iff_not_imp_left, not_mem_iff_subset hx hy] + exact id + +theorem subset_total (hx : x.IsOrdinal) (hy : y.IsOrdinal) : x ⊆ y ∨ y ⊆ x := by + obtain h | h := mem_or_subset hx hy + · exact Or.inl (hy.subset_of_mem h) + · exact Or.inr h + +theorem mem_trichotomous (hx : x.IsOrdinal) (hy : y.IsOrdinal) : x ∈ y ∨ x = y ∨ y ∈ x := by + rw [eq_comm, ← subset_iff_eq_or_mem hy hx] + exact mem_or_subset hx hy + +protected theorem isTrichotomous (h : x.IsOrdinal) : IsTrichotomous x.toSet (Subrel (· ∈ ·) _) := + ⟨fun ⟨a, ha⟩ ⟨b, hb⟩ ↦ by simpa using mem_trichotomous (h.mem ha) (h.mem hb)⟩ + +/-- An ordinal is a transitive set, trichotomous under membership. -/ +theorem _root_.ZFSet.isOrdinal_iff_isTrichotomous : + x.IsOrdinal ↔ x.IsTransitive ∧ IsTrichotomous x.toSet (Subrel (· ∈ ·) _) where + mp h := ⟨h.isTransitive, h.isTrichotomous⟩ + mpr := by + rintro ⟨h₁, h₂⟩ + rw [isOrdinal_iff_isTrans] + refine ⟨h₁, ⟨@fun y z w hyz hzw ↦ ?_⟩⟩ + obtain hyw | rfl | hwy := trichotomous_of (Subrel (· ∈ ·) _) y w + · exact hyw + · cases asymm hyz hzw + · cases mem_wf.asymmetric₃ _ _ _ hyz hzw hwy + +protected theorem isWellOrder (h : x.IsOrdinal) : IsWellOrder x.toSet (Subrel (· ∈ ·) _) where + wf := (Subrel.relEmbedding _ _).wellFounded mem_wf + trans := h.isTrans.1 + trichotomous := h.isTrichotomous.1 + +/-- An ordinal is a transitive set, well-ordered under membership. -/ +theorem _root_.ZFSet.isOrdinal_iff_isWellOrder : x.IsOrdinal ↔ + x.IsTransitive ∧ IsWellOrder x.toSet (Subrel (· ∈ ·) _) := by + use fun h ↦ ⟨h.isTransitive, h.isWellOrder⟩ + rintro ⟨h₁, h₂⟩ + refine isOrdinal_iff_isTrans.2 ⟨h₁, ?_⟩ + infer_instance + +end IsOrdinal + +@[simp] +theorem isOrdinal_empty : IsOrdinal ∅ := + ⟨isTransitive_empty, fun _ _ H ↦ (not_mem_empty _ H).elim⟩ + +/-- The **Burali-Forti paradox**: ordinals form a proper class. -/ +theorem isOrdinal_not_mem_univ : IsOrdinal ∉ Class.univ.{u} := by + rintro ⟨x, hx, -⟩ + suffices IsOrdinal x by + apply Class.mem_irrefl x + rwa [Class.coe_mem, hx] + refine ⟨fun y hy z hz ↦ ?_, fun hyz hzw hwx ↦ ?_⟩ <;> rw [← Class.coe_apply, hx] at * + exacts [hy.mem hz, hwx.mem_trans hyz hzw] end ZFSet From 7ce82ed554c5bea45b9e1a8fa13ba377b6de0e52 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sun, 15 Dec 2024 01:34:25 +0000 Subject: [PATCH 723/829] feat(SetTheory/ZFC/Basic): generalize universes of `range` (#17016) --- Mathlib/SetTheory/ZFC/Basic.lean | 19 +++++++++---------- Mathlib/SetTheory/ZFC/Rank.lean | 3 ++- 2 files changed, 11 insertions(+), 11 deletions(-) diff --git a/Mathlib/SetTheory/ZFC/Basic.lean b/Mathlib/SetTheory/ZFC/Basic.lean index 7f81281646749..c590a930ecd31 100644 --- a/Mathlib/SetTheory/ZFC/Basic.lean +++ b/Mathlib/SetTheory/ZFC/Basic.lean @@ -1209,25 +1209,24 @@ theorem toSet_image (f : ZFSet → ZFSet) [Definable₁ f] (x : ZFSet) : ext simp -/-- The range of an indexed family of sets. The universes allow for a more general index type - without manual use of `ULift`. -/ -noncomputable def range {α : Type u} (f : α → ZFSet.{max u v}) : ZFSet.{max u v} := - ⟦⟨ULift.{v} α, Quotient.out ∘ f ∘ ULift.down⟩⟧ +/-- The range of a type-indexed family of sets. -/ +noncomputable def range {α} [Small.{u} α] (f : α → ZFSet.{u}) : ZFSet.{u} := + ⟦⟨_, Quotient.out ∘ f ∘ (equivShrink α).symm⟩⟧ @[simp] -theorem mem_range {α : Type u} {f : α → ZFSet.{max u v}} {x : ZFSet.{max u v}} : - x ∈ range.{u, v} f ↔ x ∈ Set.range f := +theorem mem_range {α} [Small.{u} α] {f : α → ZFSet.{u}} {x : ZFSet.{u}} : + x ∈ range f ↔ x ∈ Set.range f := Quotient.inductionOn x fun y => by constructor · rintro ⟨z, hz⟩ - exact ⟨z.down, Quotient.eq_mk_iff_out.2 hz.symm⟩ + exact ⟨(equivShrink α).symm z, Quotient.eq_mk_iff_out.2 hz.symm⟩ · rintro ⟨z, hz⟩ - use ULift.up z + use equivShrink α z simpa [hz] using PSet.Equiv.symm (Quotient.mk_out y) @[simp] -theorem toSet_range {α : Type u} (f : α → ZFSet.{max u v}) : - (range.{u, v} f).toSet = Set.range f := by +theorem toSet_range {α} [Small.{u} α] (f : α → ZFSet.{u}) : + (range f).toSet = Set.range f := by ext simp diff --git a/Mathlib/SetTheory/ZFC/Rank.lean b/Mathlib/SetTheory/ZFC/Rank.lean index 0aaa4d9afbda4..0f0e5aeafd4a9 100644 --- a/Mathlib/SetTheory/ZFC/Rank.lean +++ b/Mathlib/SetTheory/ZFC/Rank.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Dexin Zhang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Dexin Zhang -/ +import Mathlib.Logic.UnivLE import Mathlib.SetTheory.Ordinal.Rank import Mathlib.SetTheory.ZFC.Basic @@ -194,7 +195,7 @@ theorem le_succ_rank_sUnion (x : ZFSet) : rank x ≤ succ (rank (⋃₀ x)) := b exists z @[simp] -theorem rank_range {α : Type u} (f : α → ZFSet.{max u v}) : +theorem rank_range {α : Type*} [Small.{u} α] (f : α → ZFSet.{u}) : rank (range f) = ⨆ i, succ (rank (f i)) := by apply (Ordinal.iSup_le _).antisymm' · simpa [rank_le_iff, ← succ_le_iff] using Ordinal.le_iSup _ From af60d5dfe21be6d33cc1592aec1b4a598efdd023 Mon Sep 17 00:00:00 2001 From: pauldavidrowe Date: Sun, 15 Dec 2024 02:03:32 +0000 Subject: [PATCH 724/829] feat: add theorems about isLUB and isGLB for Finset.sup and Finset.inf (#19961) feat: add theorems regarding isLUB and isGLB for Finset.sup and Finset.inf theorems were previously missing stating that `Finset.sup` produces a least upper bound and `Finset.inf` produces a greatest lower bound. This commit adds those theorems as well as the corresponding one for `Finset.sup'` and `Finset.inf'`. --- Mathlib/Data/Finset/Lattice/Fold.lean | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/Mathlib/Data/Finset/Lattice/Fold.lean b/Mathlib/Data/Finset/Lattice/Fold.lean index 506bb676f71b9..21cfa817f5ec8 100644 --- a/Mathlib/Data/Finset/Lattice/Fold.lean +++ b/Mathlib/Data/Finset/Lattice/Fold.lean @@ -106,6 +106,9 @@ theorem sup_const_le : (s.sup fun _ => a) ≤ a := theorem le_sup {b : β} (hb : b ∈ s) : f b ≤ s.sup f := Finset.sup_le_iff.1 le_rfl _ hb +theorem isLUB_sup (s : Finset α) : IsLUB s (sup s id) := + ⟨fun x h => id_eq x ▸ le_sup h, fun _ h => Finset.sup_le h⟩ + theorem le_sup_of_le {b : β} (hb : b ∈ s) (h : a ≤ f b) : a ≤ s.sup f := h.trans <| le_sup hb theorem sup_union [DecidableEq β] : (s₁ ∪ s₂).sup f = s₁.sup f ⊔ s₂.sup f := @@ -337,6 +340,9 @@ theorem le_inf_const_le : a ≤ s.inf fun _ => a := theorem inf_le {b : β} (hb : b ∈ s) : s.inf f ≤ f b := Finset.le_inf_iff.1 le_rfl _ hb +theorem isGLB_inf (s : Finset α) : IsGLB s (inf s id) := + ⟨fun x h => id_eq x ▸ inf_le h, fun _ h => Finset.le_inf h⟩ + theorem inf_le_of_le {b : β} (hb : b ∈ s) (h : f b ≤ a) : s.inf f ≤ a := (inf_le hb).trans h theorem inf_union [DecidableEq β] : (s₁ ∪ s₂).inf f = s₁.inf f ⊓ s₂.inf f := @@ -746,6 +752,10 @@ alias ⟨_, sup'_le⟩ := sup'_le_iff theorem le_sup' {b : β} (h : b ∈ s) : f b ≤ s.sup' ⟨b, h⟩ f := (sup'_le_iff ⟨b, h⟩ f).1 le_rfl b h +set_option linter.docPrime false in +theorem isLUB_sup' {s : Finset α} (hs : s.Nonempty) : IsLUB s (sup' s hs id) := + ⟨fun x h => id_eq x ▸ le_sup' id h, fun _ h => Finset.sup'_le hs id h⟩ + theorem le_sup'_of_le {a : α} {b : β} (hb : b ∈ s) (h : a ≤ f b) : a ≤ s.sup' ⟨b, hb⟩ f := h.trans <| le_sup' _ hb @@ -916,6 +926,10 @@ theorem le_inf' {a : α} (hs : ∀ b ∈ s, a ≤ f b) : a ≤ s.inf' H f := theorem inf'_le {b : β} (h : b ∈ s) : s.inf' ⟨b, h⟩ f ≤ f b := le_sup' (α := αᵒᵈ) f h +set_option linter.docPrime false in +theorem isGLB_inf' {s : Finset α} (hs : s.Nonempty) : IsGLB s (inf' s hs id) := + ⟨fun x h => id_eq x ▸ inf'_le id h, fun _ h => Finset.le_inf' hs id h⟩ + theorem inf'_le_of_le {a : α} {b : β} (hb : b ∈ s) (h : f b ≤ a) : s.inf' ⟨b, hb⟩ f ≤ a := (inf'_le _ hb).trans h From 07a5460913024dd3de0201e48e227f5e9f5c843d Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sun, 15 Dec 2024 09:21:02 +0000 Subject: [PATCH 725/829] chore: rename Mathlib.Vector to List.Vector (#19930) Per @digama0's suggestion on [zulip](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/What.20is.20the.20plan.20for.20.60Mathlib.2EVector.60.3F/near/488525657). --- .github/PULL_REQUEST_TEMPLATE.md | 2 +- Mathlib/Computability/Halting.lean | 28 ++++++++-------- Mathlib/Computability/Partrec.lean | 20 +++++------ Mathlib/Computability/Primrec.lean | 42 ++++++++++++------------ Mathlib/Computability/TMToPartrec.lean | 20 +++++------ Mathlib/Computability/TuringMachine.lean | 10 +++--- Mathlib/Data/Finite/Vector.lean | 2 +- Mathlib/Data/Fintype/BigOperators.lean | 2 +- Mathlib/Data/Fintype/Fin.lean | 4 +-- Mathlib/Data/Fintype/Vector.lean | 4 +-- Mathlib/Data/Num/Bitwise.lean | 4 +-- Mathlib/Data/Set/Finite/List.lean | 2 +- Mathlib/Data/Sym/Basic.lean | 16 ++++----- Mathlib/Data/Sym/Sym2.lean | 6 ++-- Mathlib/Data/Vector/Basic.lean | 12 +++---- Mathlib/Data/Vector/Defs.lean | 11 +++---- Mathlib/Data/Vector/MapLemmas.lean | 5 ++- Mathlib/Data/Vector/Mem.lean | 4 +-- Mathlib/Data/Vector/Snoc.lean | 4 +-- Mathlib/Data/Vector/Zip.lean | 4 +-- Mathlib/GroupTheory/Perm/Cycle/Type.lean | 12 +++---- Mathlib/Logic/Equiv/List.lean | 6 ++-- Mathlib/Logic/Small/List.lean | 4 +-- Mathlib/SetTheory/Cardinal/Basic.lean | 6 ++-- Mathlib/Topology/List.lean | 22 ++++++------- MathlibTest/NthRewrite.lean | 2 +- MathlibTest/cc.lean | 10 +++--- scripts/nolints_prime_decls.txt | 4 +-- 28 files changed, 131 insertions(+), 137 deletions(-) diff --git a/.github/PULL_REQUEST_TEMPLATE.md b/.github/PULL_REQUEST_TEMPLATE.md index 6df9d358c6266..502f31d7d75ba 100644 --- a/.github/PULL_REQUEST_TEMPLATE.md +++ b/.github/PULL_REQUEST_TEMPLATE.md @@ -20,7 +20,7 @@ If you are moving or deleting declarations, please include these lines at the bo (that is, before the `---`) using the following format: Moves: -- Vector.* -> Mathlib.Vector.* +- Vector.* -> List.Vector.* - ... Deletions: diff --git a/Mathlib/Computability/Halting.lean b/Mathlib/Computability/Halting.lean index 7437dbc1b31ac..840833a0d356f 100644 --- a/Mathlib/Computability/Halting.lean +++ b/Mathlib/Computability/Halting.lean @@ -16,7 +16,7 @@ A universal partial recursive function, Rice's theorem, and the halting problem. * [Mario Carneiro, *Formalizing computability theory via partial recursive functions*][carneiro2019] -/ -open Mathlib (Vector) +open List (Vector) open Encodable Denumerable namespace Nat.Partrec @@ -267,19 +267,19 @@ namespace Nat open Vector Part /-- A simplified basis for `Partrec`. -/ -inductive Partrec' : ∀ {n}, (Mathlib.Vector ℕ n →. ℕ) → Prop +inductive Partrec' : ∀ {n}, (List.Vector ℕ n →. ℕ) → Prop | prim {n f} : @Primrec' n f → @Partrec' n f - | comp {m n f} (g : Fin n → Mathlib.Vector ℕ m →. ℕ) : + | comp {m n f} (g : Fin n → List.Vector ℕ m →. ℕ) : Partrec' f → (∀ i, Partrec' (g i)) → - Partrec' fun v => (Mathlib.Vector.mOfFn fun i => g i v) >>= f - | rfind {n} {f : Mathlib.Vector ℕ (n + 1) → ℕ} : + Partrec' fun v => (List.Vector.mOfFn fun i => g i v) >>= f + | rfind {n} {f : List.Vector ℕ (n + 1) → ℕ} : @Partrec' (n + 1) f → Partrec' fun v => rfind fun n => some (f (n ::ᵥ v) = 0) end Nat namespace Nat.Partrec' -open Mathlib.Vector Partrec Computable +open List.Vector Partrec Computable open Nat.Partrec' @@ -294,11 +294,11 @@ theorem to_part {n f} (pf : @Partrec' n f) : _root_.Partrec f := by this).to₂.partrec₂ exact _root_.Partrec.rfind this -theorem of_eq {n} {f g : Mathlib.Vector ℕ n →. ℕ} (hf : Partrec' f) (H : ∀ i, f i = g i) : +theorem of_eq {n} {f g : List.Vector ℕ n →. ℕ} (hf : Partrec' f) (H : ∀ i, f i = g i) : Partrec' g := (funext H : f = g) ▸ hf -theorem of_prim {n} {f : Mathlib.Vector ℕ n → ℕ} (hf : Primrec f) : @Partrec' n f := +theorem of_prim {n} {f : List.Vector ℕ n → ℕ} (hf : Primrec f) : @Partrec' n f := prim (Nat.Primrec'.of_prim hf) theorem head {n : ℕ} : @Partrec' n.succ (@head ℕ n) := @@ -315,20 +315,20 @@ protected theorem bind {n f g} (hf : @Partrec' n f) (hg : @Partrec' (n + 1) g) : exact prim (Nat.Primrec'.get _)).of_eq fun v => by simp [mOfFn, Part.bind_assoc, pure] -protected theorem map {n f} {g : Mathlib.Vector ℕ (n + 1) → ℕ} (hf : @Partrec' n f) +protected theorem map {n f} {g : List.Vector ℕ (n + 1) → ℕ} (hf : @Partrec' n f) (hg : @Partrec' (n + 1) g) : @Partrec' n fun v => (f v).map fun a => g (a ::ᵥ v) := by simpa [(Part.bind_some_eq_map _ _).symm] using hf.bind hg /-- Analogous to `Nat.Partrec'` for `ℕ`-valued functions, a predicate for partial recursive vector-valued functions. -/ -def Vec {n m} (f : Mathlib.Vector ℕ n → Mathlib.Vector ℕ m) := +def Vec {n m} (f : List.Vector ℕ n → List.Vector ℕ m) := ∀ i, Partrec' fun v => (f v).get i nonrec theorem Vec.prim {n m f} (hf : @Nat.Primrec'.Vec n m f) : Vec f := fun i => prim <| hf i protected theorem nil {n} : @Vec n 0 fun _ => nil := fun i => i.elim0 -protected theorem cons {n m} {f : Mathlib.Vector ℕ n → ℕ} {g} (hf : @Partrec' n f) +protected theorem cons {n m} {f : List.Vector ℕ n → ℕ} {g} (hf : @Partrec' n f) (hg : @Vec n m g) : Vec fun v => f v ::ᵥ g v := fun i => Fin.cases (by simpa using hf) (fun i => by simp only [hg i, get_cons_succ]) i @@ -338,11 +338,11 @@ theorem idv {n} : @Vec n n id := theorem comp' {n m f g} (hf : @Partrec' m f) (hg : @Vec n m g) : Partrec' fun v => f (g v) := (hf.comp _ hg).of_eq fun v => by simp -theorem comp₁ {n} (f : ℕ →. ℕ) {g : Mathlib.Vector ℕ n → ℕ} (hf : @Partrec' 1 fun v => f v.head) +theorem comp₁ {n} (f : ℕ →. ℕ) {g : List.Vector ℕ n → ℕ} (hf : @Partrec' 1 fun v => f v.head) (hg : @Partrec' n g) : @Partrec' n fun v => f (g v) := by simpa using hf.comp' (Partrec'.cons hg Partrec'.nil) -theorem rfindOpt {n} {f : Mathlib.Vector ℕ (n + 1) → ℕ} (hf : @Partrec' (n + 1) f) : +theorem rfindOpt {n} {f : List.Vector ℕ (n + 1) → ℕ} (hf : @Partrec' (n + 1) f) : @Partrec' n fun v => Nat.rfindOpt fun a => ofNat (Option ℕ) (f (a ::ᵥ v)) := ((rfind <| (of_prim (Primrec.nat_sub.comp (_root_.Primrec.const 1) Primrec.vector_head)).comp₁ @@ -369,7 +369,7 @@ open Nat.Partrec.Code theorem of_part : ∀ {n f}, _root_.Partrec f → @Partrec' n f := @(suffices ∀ f, Nat.Partrec f → @Partrec' 1 fun v => f v.head from fun {n f} hf => by let g := fun n₁ => - (Part.ofOption (decode (α := Mathlib.Vector ℕ n) n₁)).bind (fun a => Part.map encode (f a)) + (Part.ofOption (decode (α := List.Vector ℕ n) n₁)).bind (fun a => Part.map encode (f a)) exact (comp₁ g (this g hf) (prim Nat.Primrec'.encode)).of_eq fun i => by dsimp only [g]; simp [encodek, Part.map_id'] diff --git a/Mathlib/Computability/Partrec.lean b/Mathlib/Computability/Partrec.lean index 3ad33d20e77c5..1068bd7cf6bf7 100644 --- a/Mathlib/Computability/Partrec.lean +++ b/Mathlib/Computability/Partrec.lean @@ -20,7 +20,7 @@ using the `Part` monad, and there is an additional operation, called * [Mario Carneiro, *Formalizing computability theory via partial recursive functions*][carneiro2019] -/ -open Mathlib (Vector) +open List (Vector) open Encodable Denumerable Part attribute [-simp] not_forall @@ -314,25 +314,25 @@ theorem list_concat : Computable₂ fun l (a : α) => l ++ [a] := theorem list_length : Computable (@List.length α) := Primrec.list_length.to_comp -theorem vector_cons {n} : Computable₂ (@Mathlib.Vector.cons α n) := +theorem vector_cons {n} : Computable₂ (@List.Vector.cons α n) := Primrec.vector_cons.to_comp -theorem vector_toList {n} : Computable (@Mathlib.Vector.toList α n) := +theorem vector_toList {n} : Computable (@List.Vector.toList α n) := Primrec.vector_toList.to_comp -theorem vector_length {n} : Computable (@Mathlib.Vector.length α n) := +theorem vector_length {n} : Computable (@List.Vector.length α n) := Primrec.vector_length.to_comp -theorem vector_head {n} : Computable (@Mathlib.Vector.head α n) := +theorem vector_head {n} : Computable (@List.Vector.head α n) := Primrec.vector_head.to_comp -theorem vector_tail {n} : Computable (@Mathlib.Vector.tail α n) := +theorem vector_tail {n} : Computable (@List.Vector.tail α n) := Primrec.vector_tail.to_comp -theorem vector_get {n} : Computable₂ (@Mathlib.Vector.get α n) := +theorem vector_get {n} : Computable₂ (@List.Vector.get α n) := Primrec.vector_get.to_comp -theorem vector_ofFn' {n} : Computable (@Mathlib.Vector.ofFn α n) := +theorem vector_ofFn' {n} : Computable (@List.Vector.ofFn α n) := Primrec.vector_ofFn'.to_comp theorem fin_app {n} : Computable₂ (@id (Fin n → σ)) := @@ -524,7 +524,7 @@ end Partrec @[simp] theorem Vector.mOfFn_part_some {α n} : ∀ f : Fin n → α, - (Mathlib.Vector.mOfFn fun i => Part.some (f i)) = Part.some (Mathlib.Vector.ofFn f) := + (List.Vector.mOfFn fun i => Part.some (f i)) = Part.some (List.Vector.ofFn f) := Vector.mOfFn_pure namespace Computable @@ -643,7 +643,7 @@ theorem list_ofFn : exact list_cons.comp (hf 0) (list_ofFn fun i => hf i.succ) theorem vector_ofFn {n} {f : Fin n → α → σ} (hf : ∀ i, Computable (f i)) : - Computable fun a => Mathlib.Vector.ofFn fun i => f i a := + Computable fun a => List.Vector.ofFn fun i => f i a := (Partrec.vector_mOfFn hf).of_eq fun a => by simp end Computable diff --git a/Mathlib/Computability/Primrec.lean b/Mathlib/Computability/Primrec.lean index 1eaa637c25e5e..c1e4369eb910c 100644 --- a/Mathlib/Computability/Primrec.lean +++ b/Mathlib/Computability/Primrec.lean @@ -27,7 +27,7 @@ for this.) * [Mario Carneiro, *Formalizing computability theory via partial recursive functions*][carneiro2019] -/ -open Mathlib (Vector) +open List (Vector) open Denumerable Encodable Function namespace Nat @@ -1101,7 +1101,7 @@ def subtype {p : α → Prop} [DecidablePred p] (hp : PrimrecPred p) : Primcodab instance fin {n} : Primcodable (Fin n) := @ofEquiv _ _ (subtype <| nat_lt.comp .id (const n)) Fin.equivSubtype -instance vector {n} : Primcodable (Mathlib.Vector α n) := +instance vector {n} : Primcodable (List.Vector α n) := subtype ((@Primrec.eq ℕ _ _).comp list_length (const _)) instance finArrow {n} : Primcodable (Fin n → α) := @@ -1184,26 +1184,26 @@ theorem fin_val {n} : Primrec (fun (i : Fin n) => (i : ℕ)) := theorem fin_succ {n} : Primrec (@Fin.succ n) := fin_val_iff.1 <| by simp [succ.comp fin_val] -theorem vector_toList {n} : Primrec (@Mathlib.Vector.toList α n) := +theorem vector_toList {n} : Primrec (@List.Vector.toList α n) := subtype_val -theorem vector_toList_iff {n} {f : α → Mathlib.Vector β n} : +theorem vector_toList_iff {n} {f : α → List.Vector β n} : (Primrec fun a => (f a).toList) ↔ Primrec f := subtype_val_iff -theorem vector_cons {n} : Primrec₂ (@Mathlib.Vector.cons α n) := +theorem vector_cons {n} : Primrec₂ (@List.Vector.cons α n) := vector_toList_iff.1 <| by simpa using list_cons.comp fst (vector_toList_iff.2 snd) -theorem vector_length {n} : Primrec (@Mathlib.Vector.length α n) := +theorem vector_length {n} : Primrec (@List.Vector.length α n) := const _ -theorem vector_head {n} : Primrec (@Mathlib.Vector.head α n) := +theorem vector_head {n} : Primrec (@List.Vector.head α n) := option_some_iff.1 <| (list_head?.comp vector_toList).of_eq fun ⟨_ :: _, _⟩ => rfl -theorem vector_tail {n} : Primrec (@Mathlib.Vector.tail α n) := +theorem vector_tail {n} : Primrec (@List.Vector.tail α n) := vector_toList_iff.1 <| (list_tail.comp vector_toList).of_eq fun ⟨l, h⟩ => by cases l <;> rfl -theorem vector_get {n} : Primrec₂ (@Mathlib.Vector.get α n) := +theorem vector_get {n} : Primrec₂ (@List.Vector.get α n) := option_some_iff.1 <| (list_get?.comp (vector_toList.comp fst) (fin_val.comp snd)).of_eq fun a => by rw [Vector.get_eq_get, ← List.get?_eq_get] @@ -1216,13 +1216,13 @@ theorem list_ofFn : simpa [List.ofFn_succ] using list_cons.comp (hf 0) (list_ofFn fun i => hf i.succ) theorem vector_ofFn {n} {f : Fin n → α → σ} (hf : ∀ i, Primrec (f i)) : - Primrec fun a => Mathlib.Vector.ofFn fun i => f i a := + Primrec fun a => List.Vector.ofFn fun i => f i a := vector_toList_iff.1 <| by simp [list_ofFn hf] -theorem vector_get' {n} : Primrec (@Mathlib.Vector.get α n) := +theorem vector_get' {n} : Primrec (@List.Vector.get α n) := of_equiv_symm -theorem vector_ofFn' {n} : Primrec (@Mathlib.Vector.ofFn α n) := +theorem vector_ofFn' {n} : Primrec (@List.Vector.ofFn α n) := of_equiv theorem fin_app {n} : Primrec₂ (@id (Fin n → σ)) := @@ -1242,30 +1242,30 @@ end Primrec namespace Nat -open Mathlib.Vector +open List.Vector /-- An alternative inductive definition of `Primrec` which does not use the pairing function on ℕ, and so has to work with n-ary functions on ℕ instead of unary functions. We prove that this is equivalent to the regular notion in `to_prim` and `of_prim`. -/ -inductive Primrec' : ∀ {n}, (Mathlib.Vector ℕ n → ℕ) → Prop +inductive Primrec' : ∀ {n}, (List.Vector ℕ n → ℕ) → Prop | zero : @Primrec' 0 fun _ => 0 | succ : @Primrec' 1 fun v => succ v.head | get {n} (i : Fin n) : Primrec' fun v => v.get i - | comp {m n f} (g : Fin n → Mathlib.Vector ℕ m → ℕ) : - Primrec' f → (∀ i, Primrec' (g i)) → Primrec' fun a => f (Mathlib.Vector.ofFn fun i => g i a) + | comp {m n f} (g : Fin n → List.Vector ℕ m → ℕ) : + Primrec' f → (∀ i, Primrec' (g i)) → Primrec' fun a => f (List.Vector.ofFn fun i => g i a) | prec {n f g} : @Primrec' n f → @Primrec' (n + 2) g → - Primrec' fun v : Mathlib.Vector ℕ (n + 1) => + Primrec' fun v : List.Vector ℕ (n + 1) => v.head.rec (f v.tail) fun y IH => g (y ::ᵥ IH ::ᵥ v.tail) end Nat namespace Nat.Primrec' -open Mathlib.Vector Primrec +open List.Vector Primrec theorem to_prim {n f} (pf : @Nat.Primrec' n f) : Primrec f := by induction pf with @@ -1281,7 +1281,7 @@ theorem to_prim {n f} (pf : @Nat.Primrec' n f) : Primrec f := by Primrec.vector_cons.comp (Primrec.snd.comp .snd) <| (@Primrec.vector_tail _ _ (n + 1)).comp .fst).to₂ -theorem of_eq {n} {f g : Mathlib.Vector ℕ n → ℕ} (hf : Primrec' f) (H : ∀ i, f i = g i) : +theorem of_eq {n} {f g : List.Vector ℕ n → ℕ} (hf : Primrec' f) (H : ∀ i, f i = g i) : Primrec' g := (funext H : f = g) ▸ hf @@ -1297,7 +1297,7 @@ theorem tail {n f} (hf : @Primrec' n f) : @Primrec' n.succ fun v => f v.tail := rw [← ofFn_get v.tail]; congr; funext i; simp /-- A function from vectors to vectors is primitive recursive when all of its projections are. -/ -def Vec {n m} (f : Mathlib.Vector ℕ n → Mathlib.Vector ℕ m) : Prop := +def Vec {n m} (f : List.Vector ℕ n → List.Vector ℕ m) : Prop := ∀ i, Primrec' fun v => (f v).get i protected theorem nil {n} : @Vec n 0 fun _ => nil := fun i => i.elim0 @@ -1392,7 +1392,7 @@ theorem unpair₂ {n f} (hf : @Primrec' n f) : @Primrec' n fun v => (f v).unpair theorem of_prim {n f} : Primrec f → @Primrec' n f := suffices ∀ f, Nat.Primrec f → @Primrec' 1 fun v => f v.head from fun hf => (pred.comp₁ _ <| - (this _ hf).comp₁ (fun m => Encodable.encode <| (@decode (Mathlib.Vector ℕ n) _ m).map f) + (this _ hf).comp₁ (fun m => Encodable.encode <| (@decode (List.Vector ℕ n) _ m).map f) Primrec'.encode).of_eq fun i => by simp [encodek] fun f hf => by diff --git a/Mathlib/Computability/TMToPartrec.lean b/Mathlib/Computability/TMToPartrec.lean index effb93335081f..8b095162d02d9 100644 --- a/Mathlib/Computability/TMToPartrec.lean +++ b/Mathlib/Computability/TMToPartrec.lean @@ -23,7 +23,7 @@ Turing machine for evaluating these functions. This amounts to a constructive pr * `PartrecToTM2.tr`: A TM2 turing machine which can evaluate `code` programs -/ -open Mathlib (Vector) +open List (Vector) open Function (update) @@ -239,14 +239,14 @@ def prec (f g : Code) : Code := attribute [-simp] Part.bind_eq_bind Part.map_eq_map Part.pure_eq_some -theorem exists_code.comp {m n} {f : Mathlib.Vector ℕ n →. ℕ} {g : Fin n → Mathlib.Vector ℕ m →. ℕ} - (hf : ∃ c : Code, ∀ v : Mathlib.Vector ℕ n, c.eval v.1 = pure <$> f v) - (hg : ∀ i, ∃ c : Code, ∀ v : Mathlib.Vector ℕ m, c.eval v.1 = pure <$> g i v) : - ∃ c : Code, ∀ v : Mathlib.Vector ℕ m, - c.eval v.1 = pure <$> ((Mathlib.Vector.mOfFn fun i => g i v) >>= f) := by +theorem exists_code.comp {m n} {f : List.Vector ℕ n →. ℕ} {g : Fin n → List.Vector ℕ m →. ℕ} + (hf : ∃ c : Code, ∀ v : List.Vector ℕ n, c.eval v.1 = pure <$> f v) + (hg : ∀ i, ∃ c : Code, ∀ v : List.Vector ℕ m, c.eval v.1 = pure <$> g i v) : + ∃ c : Code, ∀ v : List.Vector ℕ m, + c.eval v.1 = pure <$> ((List.Vector.mOfFn fun i => g i v) >>= f) := by rsuffices ⟨cg, hg⟩ : - ∃ c : Code, ∀ v : Mathlib.Vector ℕ m, - c.eval v.1 = Subtype.val <$> Mathlib.Vector.mOfFn fun i => g i v + ∃ c : Code, ∀ v : List.Vector ℕ m, + c.eval v.1 = Subtype.val <$> List.Vector.mOfFn fun i => g i v · obtain ⟨cf, hf⟩ := hf exact ⟨cf.comp cg, fun v => by @@ -261,8 +261,8 @@ theorem exists_code.comp {m n} {f : Mathlib.Vector ℕ n →. ℕ} {g : Fin n simp [Vector.mOfFn, hg₁, map_bind, seq_bind_eq, bind_assoc, (· ∘ ·), hl] rfl⟩ -theorem exists_code {n} {f : Mathlib.Vector ℕ n →. ℕ} (hf : Nat.Partrec' f) : - ∃ c : Code, ∀ v : Mathlib.Vector ℕ n, c.eval v.1 = pure <$> f v := by +theorem exists_code {n} {f : List.Vector ℕ n →. ℕ} (hf : Nat.Partrec' f) : + ∃ c : Code, ∀ v : List.Vector ℕ n, c.eval v.1 = pure <$> f v := by induction hf with | prim hf => induction hf with diff --git a/Mathlib/Computability/TuringMachine.lean b/Mathlib/Computability/TuringMachine.lean index 0f1145c0c16ee..3af2e1dc86a62 100644 --- a/Mathlib/Computability/TuringMachine.lean +++ b/Mathlib/Computability/TuringMachine.lean @@ -60,7 +60,7 @@ Given these parameters, there are a few common structures for the model that ari assert_not_exists MonoidWithZero -open Mathlib (Vector) +open List (Vector) open Relation open Nat (iterate) @@ -1443,7 +1443,7 @@ open TM1 variable {Γ : Type*} theorem exists_enc_dec [Inhabited Γ] [Finite Γ] : - ∃ (n : ℕ) (enc : Γ → Mathlib.Vector Bool n) (dec : Mathlib.Vector Bool n → Γ), + ∃ (n : ℕ) (enc : Γ → List.Vector Bool n) (dec : List.Vector Bool n → Γ), enc default = Vector.replicate n false ∧ ∀ a, dec (enc a) = a := by rcases Finite.exists_equiv_fin Γ with ⟨n, ⟨e⟩⟩ letI : DecidableEq Γ := e.decidableEq @@ -1475,13 +1475,13 @@ local notation "Stmt'₁" => Stmt Bool Λ'₁ σ local notation "Cfg'₁" => Cfg Bool Λ'₁ σ /-- Read a vector of length `n` from the tape. -/ -def readAux : ∀ n, (Mathlib.Vector Bool n → Stmt'₁) → Stmt'₁ +def readAux : ∀ n, (List.Vector Bool n → Stmt'₁) → Stmt'₁ | 0, f => f Vector.nil | i + 1, f => Stmt.branch (fun a _ ↦ a) (Stmt.move Dir.right <| readAux i fun v ↦ f (true ::ᵥ v)) (Stmt.move Dir.right <| readAux i fun v ↦ f (false ::ᵥ v)) -variable {n : ℕ} (enc : Γ → Mathlib.Vector Bool n) (dec : Mathlib.Vector Bool n → Γ) +variable {n : ℕ} (enc : Γ → List.Vector Bool n) (dec : List.Vector Bool n → Γ) /-- A move left or right corresponds to `n` moves across the super-cell. -/ def move (d : Dir) (q : Stmt'₁) : Stmt'₁ := @@ -1530,7 +1530,7 @@ theorem supportsStmt_write {S : Finset Λ'₁} {l : List Bool} {q : Stmt'₁} : theorem supportsStmt_read {S : Finset Λ'₁} : ∀ {f : Γ → Stmt'₁}, (∀ a, SupportsStmt S (f a)) → SupportsStmt S (read dec f) := suffices - ∀ (i) (f : Mathlib.Vector Bool i → Stmt'₁), + ∀ (i) (f : List.Vector Bool i → Stmt'₁), (∀ v, SupportsStmt S (f v)) → SupportsStmt S (readAux i f) from fun hf ↦ this n _ (by intro; simp only [supportsStmt_move, hf]) fun i f hf ↦ by diff --git a/Mathlib/Data/Finite/Vector.lean b/Mathlib/Data/Finite/Vector.lean index 9beba5adef683..080c6889787cd 100644 --- a/Mathlib/Data/Finite/Vector.lean +++ b/Mathlib/Data/Finite/Vector.lean @@ -10,6 +10,6 @@ import Mathlib.Data.Fintype.Card # Finiteness of vector types -/ -instance Mathlib.Vector.finite {α : Type*} [Finite α] {n : ℕ} : Finite (Vector α n) := by +instance List.Vector.finite {α : Type*} [Finite α] {n : ℕ} : Finite (Vector α n) := by haveI := Fintype.ofFinite α infer_instance diff --git a/Mathlib/Data/Fintype/BigOperators.lean b/Mathlib/Data/Fintype/BigOperators.lean index 0825cf4335415..a89cc163bdcce 100644 --- a/Mathlib/Data/Fintype/BigOperators.lean +++ b/Mathlib/Data/Fintype/BigOperators.lean @@ -173,7 +173,7 @@ theorem Fintype.card_fun [DecidableEq α] [Fintype α] [Fintype β] : @[simp] theorem card_vector [Fintype α] (n : ℕ) : - Fintype.card (Mathlib.Vector α n) = Fintype.card α ^ n := by + Fintype.card (List.Vector α n) = Fintype.card α ^ n := by rw [Fintype.ofEquiv_card]; simp /-- It is equivalent to compute the product of a function over `Fin n` or `Finset.range n`. -/ diff --git a/Mathlib/Data/Fintype/Fin.lean b/Mathlib/Data/Fintype/Fin.lean index 96b7ea32bd858..ed8d702c376fe 100644 --- a/Mathlib/Data/Fintype/Fin.lean +++ b/Mathlib/Data/Fintype/Fin.lean @@ -13,7 +13,7 @@ This file contains some basic results about the `Fintype` instance for `Fin`, especially properties of `Finset.univ : Finset (Fin n)`. -/ -open Mathlib +open List (Vector) open Finset @@ -61,7 +61,7 @@ theorem card_filter_univ_succ' (p : Fin (n + 1) → Prop) [DecidablePred p] : #{x | p x} = ite (p 0) 1 0 + #{x | p (.succ x)}:= by rw [card_filter_univ_succ]; split_ifs <;> simp [add_comm] -theorem card_filter_univ_eq_vector_get_eq_count [DecidableEq α] (a : α) (v : Mathlib.Vector α n) : +theorem card_filter_univ_eq_vector_get_eq_count [DecidableEq α] (a : α) (v : List.Vector α n) : #{i | v.get i = a} = v.toList.count a := by induction' v with n x xs hxs · simp diff --git a/Mathlib/Data/Fintype/Vector.lean b/Mathlib/Data/Fintype/Vector.lean index 9b833c039d387..ca742df20f51a 100644 --- a/Mathlib/Data/Fintype/Vector.lean +++ b/Mathlib/Data/Fintype/Vector.lean @@ -10,11 +10,11 @@ import Mathlib.Data.Sym.Basic # `Vector α n` and `Sym α n` are fintypes when `α` is. -/ -open Mathlib (Vector) +open List (Vector) variable {α : Type*} -instance Vector.fintype [Fintype α] {n : ℕ} : Fintype (Mathlib.Vector α n) := +instance Vector.fintype [Fintype α] {n : ℕ} : Fintype (List.Vector α n) := Fintype.ofEquiv _ (Equiv.vectorEquivFin _ _).symm instance [DecidableEq α] [Fintype α] {n : ℕ} : Fintype (Sym.Sym' α n) := by diff --git a/Mathlib/Data/Num/Bitwise.lean b/Mathlib/Data/Num/Bitwise.lean index 57703e62a4600..f3a69144970a6 100644 --- a/Mathlib/Data/Num/Bitwise.lean +++ b/Mathlib/Data/Num/Bitwise.lean @@ -16,7 +16,7 @@ import Mathlib.Data.Vector.Basic * arithmetic operations for `SNum`. -/ -open Mathlib +open List (Vector) namespace PosNum @@ -409,7 +409,7 @@ end SNum namespace SNum /-- `a.bits n` is the vector of the `n` first bits of `a` (starting from the LSB). -/ -def bits : SNum → ∀ n, Mathlib.Vector Bool n +def bits : SNum → ∀ n, List.Vector Bool n | _, 0 => Vector.nil | p, n + 1 => head p ::ᵥ bits (tail p) n diff --git a/Mathlib/Data/Set/Finite/List.lean b/Mathlib/Data/Set/Finite/List.lean index 1282e6b3c5da0..2b630c4dfa412 100644 --- a/Mathlib/Data/Set/Finite/List.lean +++ b/Mathlib/Data/Set/Finite/List.lean @@ -24,7 +24,7 @@ assert_not_exists MonoidWithZero namespace List variable (α : Type*) [Finite α] (n : ℕ) -lemma finite_length_eq : {l : List α | l.length = n}.Finite := Mathlib.Vector.finite +lemma finite_length_eq : {l : List α | l.length = n}.Finite := List.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 diff --git a/Mathlib/Data/Sym/Basic.lean b/Mathlib/Data/Sym/Basic.lean index 4daed5e6c93fa..2733966c42fc7 100644 --- a/Mathlib/Data/Sym/Basic.lean +++ b/Mathlib/Data/Sym/Basic.lean @@ -28,7 +28,7 @@ symmetric powers -/ assert_not_exists MonoidWithZero -open Mathlib (Vector) +open List (Vector) open Function /-- The nth symmetric power is n-tuples up to permutation. We define it @@ -55,7 +55,7 @@ instance {α : Type*} {n : ℕ} [DecidableEq α] : DecidableEq (Sym α n) := See note [reducible non-instances]. -/ -abbrev Mathlib.Vector.Perm.isSetoid (α : Type*) (n : ℕ) : Setoid (Vector α n) := +abbrev List.Vector.Perm.isSetoid (α : Type*) (n : ℕ) : Setoid (Vector α n) := (List.isSetoid α).comap Subtype.val attribute [local instance] Vector.Perm.isSetoid @@ -119,20 +119,20 @@ theorem coe_cons (s : Sym α n) (a : α) : (a ::ₛ s : Multiset α) = a ::ₘ s /-- This is the quotient map that takes a list of n elements as an n-tuple and produces an nth symmetric power. -/ -def ofVector : Mathlib.Vector α n → Sym α n := +def ofVector : List.Vector α n → Sym α n := fun x => ⟨↑x.val, (Multiset.coe_card _).trans x.2⟩ /-- This is the quotient map that takes a list of n elements as an n-tuple and produces an nth symmetric power. -/ -instance : Coe (Mathlib.Vector α n) (Sym α n) where coe x := ofVector x +instance : Coe (List.Vector α n) (Sym α n) where coe x := ofVector x @[simp] -theorem ofVector_nil : ↑(Vector.nil : Mathlib.Vector α 0) = (Sym.nil : Sym α 0) := +theorem ofVector_nil : ↑(Vector.nil : List.Vector α 0) = (Sym.nil : Sym α 0) := rfl @[simp] -theorem ofVector_cons (a : α) (v : Mathlib.Vector α n) : +theorem ofVector_cons (a : α) (v : List.Vector α n) : ↑(Vector.cons a v) = a ::ₛ (↑v : Sym α n) := by cases v rfl @@ -180,13 +180,13 @@ theorem mem_cons_of_mem (h : a ∈ s) : a ∈ b ::ₛ s := theorem mem_cons_self (a : α) (s : Sym α n) : a ∈ a ::ₛ s := Multiset.mem_cons_self a s.1 -theorem cons_of_coe_eq (a : α) (v : Mathlib.Vector α n) : a ::ₛ (↑v : Sym α n) = ↑(a ::ᵥ v) := +theorem cons_of_coe_eq (a : α) (v : List.Vector α n) : a ::ₛ (↑v : Sym α n) = ↑(a ::ᵥ v) := Subtype.ext <| by cases v rfl open scoped List in -theorem sound {a b : Mathlib.Vector α n} (h : a.val ~ b.val) : (↑a : Sym α n) = ↑b := +theorem sound {a b : List.Vector α n} (h : a.val ~ b.val) : (↑a : Sym α n) = ↑b := Subtype.ext <| Quotient.sound h /-- `erase s a h` is the sym that subtracts 1 from the diff --git a/Mathlib/Data/Sym/Sym2.lean b/Mathlib/Data/Sym/Sym2.lean index a827017219d56..3b6df7f76fa41 100644 --- a/Mathlib/Data/Sym/Sym2.lean +++ b/Mathlib/Data/Sym/Sym2.lean @@ -44,7 +44,7 @@ symmetric square, unordered pairs, symmetric powers assert_not_exists MonoidWithZero -open Mathlib (Vector) +open List (Vector) open Finset Function Sym universe u @@ -584,9 +584,9 @@ section SymEquiv /-! ### Equivalence to the second symmetric power -/ -attribute [local instance] Vector.Perm.isSetoid +attribute [local instance] List.Vector.Perm.isSetoid -private def fromVector : Mathlib.Vector α 2 → α × α +private def fromVector : List.Vector α 2 → α × α | ⟨[a, b], _⟩ => (a, b) private theorem perm_card_two_iff {a₁ b₁ a₂ b₂ : α} : diff --git a/Mathlib/Data/Vector/Basic.lean b/Mathlib/Data/Vector/Basic.lean index f965bcd759bc9..5277e338c4169 100644 --- a/Mathlib/Data/Vector/Basic.lean +++ b/Mathlib/Data/Vector/Basic.lean @@ -20,9 +20,7 @@ universe u variable {α β γ σ φ : Type*} {m n : ℕ} -namespace Mathlib - -namespace Vector +namespace List.Vector @[inherit_doc] infixr:67 " ::ᵥ " => Vector.cons @@ -622,7 +620,9 @@ theorem prod_set [Monoid α] (v : Vector α n) (i : Fin n) (a : α) : refine (List.prod_set v.toList i a).trans ?_ simp_all -@[to_additive] +/-- Variant of `List.Vector.prod_set` that multiplies by the inverse of the replaced element.-/ +@[to_additive + "Variant of `List.Vector.sum_set` that subtracts the inverse of the replaced element."] theorem prod_set' [CommGroup α] (v : Vector α n) (i : Fin n) (a : α) : (v.set i a).toList.prod = v.toList.prod * (v.get i)⁻¹ * a := by refine (List.prod_set' v.toList i a).trans ?_ @@ -790,6 +790,4 @@ theorem mapAccumr₂_cons {f : α → β → σ → σ × φ} : end Simp -end Vector - -end Mathlib +end List.Vector diff --git a/Mathlib/Data/Vector/Defs.lean b/Mathlib/Data/Vector/Defs.lean index 8047a0ccff22a..db88601c6603b 100644 --- a/Mathlib/Data/Vector/Defs.lean +++ b/Mathlib/Data/Vector/Defs.lean @@ -7,18 +7,17 @@ import Mathlib.Data.List.Defs import Mathlib.Tactic.Common /-! -The type `Vector` represents lists with fixed length. +The type `List.Vector` represents lists with fixed length. -/ assert_not_exists Monoid -namespace Mathlib universe u v w /-- `Vector α n` is the type of lists of length `n` with elements of type `α`. -/ -def Vector (α : Type u) (n : ℕ) := +def List.Vector (α : Type u) (n : ℕ) := { l : List α // l.length = n } -namespace Vector +namespace List.Vector variable {α β σ φ : Type*} {n : ℕ} {p : α → Prop} @@ -240,6 +239,4 @@ lemma getElem_def (v : Vector α n) (i : ℕ) {hi : i < n} : lemma toList_getElem (v : Vector α n) (i : ℕ) {hi : i < v.toList.length} : v.toList[i] = v[i]'(by simp_all) := rfl -end Vector - -end Mathlib +end List.Vector diff --git a/Mathlib/Data/Vector/MapLemmas.lean b/Mathlib/Data/Vector/MapLemmas.lean index 63298aef59770..2d64ed93dbcf2 100644 --- a/Mathlib/Data/Vector/MapLemmas.lean +++ b/Mathlib/Data/Vector/MapLemmas.lean @@ -12,8 +12,7 @@ import Mathlib.Data.Vector.Snoc variable {α β γ ζ σ σ₁ σ₂ φ : Type*} {n : ℕ} {s : σ} {s₁ : σ₁} {s₂ : σ₂} -namespace Mathlib - +namespace List namespace Vector /-! @@ -402,4 +401,4 @@ end Flip end Vector -end Mathlib +end List diff --git a/Mathlib/Data/Vector/Mem.lean b/Mathlib/Data/Vector/Mem.lean index dc09c5bb58415..d02a37d1faaf7 100644 --- a/Mathlib/Data/Vector/Mem.lean +++ b/Mathlib/Data/Vector/Mem.lean @@ -15,7 +15,7 @@ In particular we can avoid some assumptions about types being `Inhabited`, and make more general statements about `head` and `tail`. -/ -namespace Mathlib +namespace List namespace Vector @@ -73,4 +73,4 @@ theorem mem_map_succ_iff (b : β) (v : Vector α (n + 1)) (f : α → β) : end Vector -end Mathlib +end List diff --git a/Mathlib/Data/Vector/Snoc.lean b/Mathlib/Data/Vector/Snoc.lean index ec5dccd5f02c9..400e6191966a6 100644 --- a/Mathlib/Data/Vector/Snoc.lean +++ b/Mathlib/Data/Vector/Snoc.lean @@ -16,7 +16,7 @@ import Mathlib.Data.Vector.Basic `snoc xs x` for its inductive case. Effectively doing induction from right-to-left -/ -namespace Mathlib +namespace List namespace Vector @@ -159,4 +159,4 @@ theorem mapAccumr₂_snoc (f : α → β → σ → σ × φ) (x : α) (y : β) end Simp end Vector -end Mathlib +end List diff --git a/Mathlib/Data/Vector/Zip.lean b/Mathlib/Data/Vector/Zip.lean index f8c69ad4814b4..9df0f622a8504 100644 --- a/Mathlib/Data/Vector/Zip.lean +++ b/Mathlib/Data/Vector/Zip.lean @@ -9,7 +9,7 @@ import Mathlib.Data.Vector.Basic # The `zipWith` operation on vectors. -/ -namespace Mathlib +namespace List namespace Vector @@ -47,4 +47,4 @@ end ZipWith end Vector -end Mathlib +end List diff --git a/Mathlib/GroupTheory/Perm/Cycle/Type.lean b/Mathlib/GroupTheory/Perm/Cycle/Type.lean index b47673b955137..491829c0b7160 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Type.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Type.lean @@ -33,7 +33,7 @@ In this file we define the cycle type of a permutation. namespace Equiv.Perm -open Mathlib (Vector) +open List (Vector) open Equiv List Multiset variable {α : Type*} [Fintype α] @@ -394,12 +394,12 @@ section Cauchy variable (G : Type*) [Group G] (n : ℕ) /-- The type of vectors with terms from `G`, length `n`, and product equal to `1:G`. -/ -def vectorsProdEqOne : Set (Mathlib.Vector G n) := +def vectorsProdEqOne : Set (List.Vector G n) := { v | v.toList.prod = 1 } namespace VectorsProdEqOne -theorem mem_iff {n : ℕ} (v : Mathlib.Vector G n) : v ∈ vectorsProdEqOne G n ↔ v.toList.prod = 1 := +theorem mem_iff {n : ℕ} (v : List.Vector G n) : v ∈ vectorsProdEqOne G n ↔ v.toList.prod = 1 := Iff.rfl theorem zero_eq : vectorsProdEqOne G 0 = {Vector.nil} := @@ -421,7 +421,7 @@ instance oneUnique : Unique (vectorsProdEqOne G 1) := by /-- Given a vector `v` of length `n`, make a vector of length `n + 1` whose product is `1`, by appending the inverse of the product of `v`. -/ @[simps] -def vectorEquiv : Mathlib.Vector G n ≃ vectorsProdEqOne G (n + 1) where +def vectorEquiv : List.Vector G n ≃ vectorsProdEqOne G (n + 1) where toFun v := ⟨v.toList.prod⁻¹ ::ᵥ v, by rw [mem_iff, Vector.toList_cons, List.prod_cons, inv_mul_cancel]⟩ invFun v := v.1.tail @@ -436,12 +436,12 @@ def vectorEquiv : Mathlib.Vector G n ≃ vectorsProdEqOne G (n + 1) where /-- Given a vector `v` of length `n` whose product is 1, make a vector of length `n - 1`, by deleting the last entry of `v`. -/ -def equivVector : ∀ n, vectorsProdEqOne G n ≃ Mathlib.Vector G (n - 1) +def equivVector : ∀ n, vectorsProdEqOne G n ≃ List.Vector G (n - 1) | 0 => (equivOfUnique (vectorsProdEqOne G 0) (vectorsProdEqOne G 1)).trans (vectorEquiv G 0).symm | (n + 1) => (vectorEquiv G n).symm instance [Fintype G] : Fintype (vectorsProdEqOne G n) := - Fintype.ofEquiv (Mathlib.Vector G (n - 1)) (equivVector G n).symm + Fintype.ofEquiv (List.Vector G (n - 1)) (equivVector G n).symm theorem card [Fintype G] : Fintype.card (vectorsProdEqOne G n) = Fintype.card G ^ (n - 1) := (Fintype.card_congr (equivVector G n)).trans (card_vector (n - 1)) diff --git a/Mathlib/Logic/Equiv/List.lean b/Mathlib/Logic/Equiv/List.lean index 6a6bb07e4ed21..7c84c312dfc46 100644 --- a/Mathlib/Logic/Equiv/List.lean +++ b/Mathlib/Logic/Equiv/List.lean @@ -14,7 +14,7 @@ This file defines some additional constructive equivalences using `Encodable` an function on `ℕ`. -/ -open Mathlib (Vector) +open List (Vector) open Nat List namespace Encodable @@ -127,11 +127,11 @@ noncomputable def _root_.Fintype.toEncodable (α : Type*) [Fintype α] : Encodab classical exact (Fintype.truncEncodable α).out /-- If `α` is encodable, then so is `Vector α n`. -/ -instance Mathlib.Vector.encodable [Encodable α] {n} : Encodable (Mathlib.Vector α n) := +instance List.Vector.encodable [Encodable α] {n} : Encodable (List.Vector α n) := Subtype.encodable /-- If `α` is countable, then so is `Vector α n`. -/ -instance Mathlib.Vector.countable [Countable α] {n} : Countable (Mathlib.Vector α n) := +instance List.Vector.countable [Countable α] {n} : Countable (List.Vector α n) := Subtype.countable /-- If `α` is encodable, then so is `Fin n → α`. -/ diff --git a/Mathlib/Logic/Small/List.lean b/Mathlib/Logic/Small/List.lean index f0706af610d10..f5a65c7096bb0 100644 --- a/Mathlib/Logic/Small/List.lean +++ b/Mathlib/Logic/Small/List.lean @@ -18,9 +18,9 @@ universe u v open Mathlib -instance smallVector {α : Type v} {n : ℕ} [Small.{u} α] : Small.{u} (Mathlib.Vector α n) := +instance smallVector {α : Type v} {n : ℕ} [Small.{u} α] : Small.{u} (List.Vector α n) := small_of_injective (Equiv.vectorEquivFin α n).injective instance smallList {α : Type v} [Small.{u} α] : Small.{u} (List α) := by - let e : (Σn, Mathlib.Vector α n) ≃ List α := Equiv.sigmaFiberEquiv List.length + let e : (Σn, List.Vector α n) ≃ List α := Equiv.sigmaFiberEquiv List.length exact small_of_surjective e.surjective diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index 12ba9cbde683e..0761c27a15594 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -84,7 +84,7 @@ Cantor's theorem, König's theorem, Konig's theorem assert_not_exists Field -open Mathlib (Vector) +open List (Vector) open Function Order Set noncomputable section @@ -1786,12 +1786,12 @@ theorem mk_plift_false : #(PLift False) = 0 := mk_eq_zero _ @[simp] -theorem mk_vector (α : Type u) (n : ℕ) : #(Mathlib.Vector α n) = #α ^ n := +theorem mk_vector (α : Type u) (n : ℕ) : #(List.Vector α n) = #α ^ n := (mk_congr (Equiv.vectorEquivFin α n)).trans <| by simp theorem mk_list_eq_sum_pow (α : Type u) : #(List α) = sum fun n : ℕ => #α ^ n := calc - #(List α) = #(Σn, Mathlib.Vector α n) := mk_congr (Equiv.sigmaFiberEquiv List.length).symm + #(List α) = #(Σn, List.Vector α n) := mk_congr (Equiv.sigmaFiberEquiv List.length).symm _ = sum fun n : ℕ => #α ^ n := by simp theorem mk_quot_le {α : Type u} {r : α → α → Prop} : #(Quot r) ≤ #α := diff --git a/Mathlib/Topology/List.lean b/Mathlib/Topology/List.lean index 4ca0d9581282a..c6bac03cf1d2e 100644 --- a/Mathlib/Topology/List.lean +++ b/Mathlib/Topology/List.lean @@ -182,18 +182,18 @@ namespace Vector open List -open Mathlib (Vector) +open List (Vector) -instance (n : ℕ) : TopologicalSpace (Mathlib.Vector α n) := by unfold Mathlib.Vector; infer_instance +instance (n : ℕ) : TopologicalSpace (List.Vector α n) := by unfold List.Vector; infer_instance -theorem tendsto_cons {n : ℕ} {a : α} {l : Mathlib.Vector α n} : - Tendsto (fun p : α × Mathlib.Vector α n => p.1 ::ᵥ p.2) (𝓝 a ×ˢ 𝓝 l) (𝓝 (a ::ᵥ l)) := by +theorem tendsto_cons {n : ℕ} {a : α} {l : List.Vector α n} : + Tendsto (fun p : α × List.Vector α n => p.1 ::ᵥ p.2) (𝓝 a ×ˢ 𝓝 l) (𝓝 (a ::ᵥ l)) := by rw [tendsto_subtype_rng, Vector.cons_val] exact tendsto_fst.cons (Tendsto.comp continuousAt_subtype_val tendsto_snd) theorem tendsto_insertIdx {n : ℕ} {i : Fin (n + 1)} {a : α} : - ∀ {l : Mathlib.Vector α n}, - Tendsto (fun p : α × Mathlib.Vector α n => Vector.insertIdx p.1 i p.2) (𝓝 a ×ˢ 𝓝 l) + ∀ {l : List.Vector α n}, + Tendsto (fun p : α × List.Vector α n => Vector.insertIdx p.1 i p.2) (𝓝 a ×ˢ 𝓝 l) (𝓝 (Vector.insertIdx a i l)) | ⟨l, hl⟩ => by rw [Vector.insertIdx, tendsto_subtype_rng] @@ -204,29 +204,29 @@ theorem tendsto_insertIdx {n : ℕ} {i : Fin (n + 1)} {a : α} : /-- Continuity of `Vector.insertIdx`. -/ theorem continuous_insertIdx' {n : ℕ} {i : Fin (n + 1)} : - Continuous fun p : α × Mathlib.Vector α n => Vector.insertIdx p.1 i p.2 := + Continuous fun p : α × List.Vector α n => Vector.insertIdx p.1 i p.2 := continuous_iff_continuousAt.mpr fun ⟨a, l⟩ => by rw [ContinuousAt, nhds_prod_eq]; exact tendsto_insertIdx @[deprecated (since := "2024-10-21")] alias continuous_insertNth' := continuous_insertIdx' -theorem continuous_insertIdx {n : ℕ} {i : Fin (n + 1)} {f : β → α} {g : β → Mathlib.Vector α n} +theorem continuous_insertIdx {n : ℕ} {i : Fin (n + 1)} {f : β → α} {g : β → List.Vector α n} (hf : Continuous f) (hg : Continuous g) : Continuous fun b => Vector.insertIdx (f b) i (g b) := continuous_insertIdx'.comp (hf.prod_mk hg : _) @[deprecated (since := "2024-10-21")] alias continuous_insertNth := continuous_insertIdx theorem continuousAt_eraseIdx {n : ℕ} {i : Fin (n + 1)} : - ∀ {l : Mathlib.Vector α (n + 1)}, ContinuousAt (Mathlib.Vector.eraseIdx i) l + ∀ {l : List.Vector α (n + 1)}, ContinuousAt (List.Vector.eraseIdx i) l | ⟨l, hl⟩ => by - rw [ContinuousAt, Mathlib.Vector.eraseIdx, tendsto_subtype_rng] + rw [ContinuousAt, List.Vector.eraseIdx, tendsto_subtype_rng] simp only [Vector.eraseIdx_val] exact Tendsto.comp List.tendsto_eraseIdx continuousAt_subtype_val @[deprecated (since := "2024-05-04")] alias continuousAt_removeNth := continuousAt_eraseIdx theorem continuous_eraseIdx {n : ℕ} {i : Fin (n + 1)} : - Continuous (Mathlib.Vector.eraseIdx i : Mathlib.Vector α (n + 1) → Mathlib.Vector α n) := + Continuous (List.Vector.eraseIdx i : List.Vector α (n + 1) → List.Vector α n) := continuous_iff_continuousAt.mpr fun ⟨_a, _l⟩ => continuousAt_eraseIdx @[deprecated (since := "2024-05-04")] alias continuous_removeNth := continuous_eraseIdx diff --git a/MathlibTest/NthRewrite.lean b/MathlibTest/NthRewrite.lean index 3b82ffdb5998a..b1c125016b3f0 100644 --- a/MathlibTest/NthRewrite.lean +++ b/MathlibTest/NthRewrite.lean @@ -16,7 +16,7 @@ example [AddZeroClass G] {a : G} : a + a = a + (a + 0) := by structure F where (a : ℕ) - (v : Mathlib.Vector ℕ a) + (v : List.Vector ℕ a) (p : v.val = []) example (f : F) : f.v.val = [] := by diff --git a/MathlibTest/cc.lean b/MathlibTest/cc.lean index 26e0d6ac83f4d..8fd7047024b7e 100644 --- a/MathlibTest/cc.lean +++ b/MathlibTest/cc.lean @@ -8,7 +8,7 @@ set_option linter.unusedVariables false section CC1 -open Mathlib (Vector) +open List (Vector) open Mathlib.Tactic.CC open Lean Meta Elab Tactic @@ -221,20 +221,20 @@ universe u open Mathlib axiom app : {α : Type u} → {n m : Nat} → - Mathlib.Vector α m → Mathlib.Vector α n → Mathlib.Vector α (m + n) + List.Vector α m → List.Vector α n → List.Vector α (m + n) example (n1 n2 n3 : Nat) - (v1 w1 : Mathlib.Vector Nat n1) (w1' : Mathlib.Vector Nat n3) (v2 w2 : Mathlib.Vector Nat n2) : + (v1 w1 : List.Vector Nat n1) (w1' : List.Vector Nat n3) (v2 w2 : List.Vector Nat n2) : n1 = n3 → v1 = w1 → HEq w1 w1' → v2 = w2 → HEq (app v1 v2) (app w1' w2) := by cc example (n1 n2 n3 : Nat) - (v1 w1 : Mathlib.Vector Nat n1) (w1' : Mathlib.Vector Nat n3) (v2 w2 : Mathlib.Vector Nat n2) : + (v1 w1 : List.Vector Nat n1) (w1' : List.Vector Nat n3) (v2 w2 : List.Vector Nat n2) : HEq n1 n3 → v1 = w1 → HEq w1 w1' → HEq v2 w2 → HEq (app v1 v2) (app w1' w2) := by cc example (n1 n2 n3 : Nat) - (v1 w1 v : Mathlib.Vector Nat n1) (w1' : Mathlib.Vector Nat n3) (v2 w2 w : Mathlib.Vector Nat n2) : + (v1 w1 v : List.Vector Nat n1) (w1' : List.Vector Nat n3) (v2 w2 w : List.Vector Nat n2) : HEq n1 n3 → v1 = w1 → HEq w1 w1' → HEq v2 w2 → HEq (app w1' w2) (app v w) → app v1 v2 = app v w := by cc diff --git a/scripts/nolints_prime_decls.txt b/scripts/nolints_prime_decls.txt index 141e187da20d1..6924760aaad11 100644 --- a/scripts/nolints_prime_decls.txt +++ b/scripts/nolints_prime_decls.txt @@ -2549,8 +2549,8 @@ Mathlib.Tactic.ComputeDegree.degree_eq_of_le_of_coeff_ne_zero' Mathlib.Tactic.Group.zpow_trick_one' Mathlib.Tactic.Ring.atom_pf' Mathlib.Util.addAndCompile' -Mathlib.Vector.eraseIdx_insertNth' -Mathlib.Vector.prod_set' +List.Vector.eraseIdx_insertNth' +List.Vector.prod_set' Mathlib.WhatsNew.mkHeader' Matrix.blockDiag'_blockDiagonal' Matrix.blockDiagonal'_apply' From ffcd1a0e57246087fd7c9636a6d83292bf3b902f Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sun, 15 Dec 2024 13:33:16 +0000 Subject: [PATCH 726/829] chore: revert removal of tactic usage in Archive (#19974) This reverts a change from #19691, which unnecessarily removed a tactic invocation in favour of writing out lemmas. The archive should be demonstrating good style, and unless there is a verbosity or speed issue, robust and well specified tactics should always be preferred over proofs that require knowing lemmas (even if, as in this case, those lemmas are easy to identify and use). --- Archive/Imo/Imo1964Q1.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Archive/Imo/Imo1964Q1.lean b/Archive/Imo/Imo1964Q1.lean index 6a3e8cba29dfd..ed44e0d656911 100644 --- a/Archive/Imo/Imo1964Q1.lean +++ b/Archive/Imo/Imo1964Q1.lean @@ -5,6 +5,7 @@ Authors: Kevin Buzzard -/ import Mathlib.Tactic.IntervalCases import Mathlib.Data.Nat.ModEq +import Mathlib.Tactic.Ring /-! # IMO 1964 Q1 @@ -25,7 +26,7 @@ theorem two_pow_mod_seven (n : ℕ) : 2 ^ n ≡ 2 ^ (n % 3) [MOD 7] := calc 2 ^ n = 2 ^ (3 * (n / 3) + t) := by rw [Nat.div_add_mod] _ = (2 ^ 3) ^ (n / 3) * 2 ^ t := by rw [pow_add, pow_mul] _ ≡ 1 ^ (n / 3) * 2 ^ t [MOD 7] := by gcongr; decide - _ = 2 ^ t := by rw [one_pow, one_mul] + _ = 2 ^ t := by ring end Imo1964Q1 From 00a4da3dfe01780cafe9613f222d77d7caff4056 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sun, 15 Dec 2024 23:28:00 +0000 Subject: [PATCH 727/829] chore(SetTheory/Ordinal/NaturalOps): address porting notes (#19237) We also private/deprecate some primed lemmas which are really only useful to prove associativity of multiplication. --- Mathlib/SetTheory/Ordinal/NaturalOps.lean | 121 +++++++++------------- 1 file changed, 48 insertions(+), 73 deletions(-) diff --git a/Mathlib/SetTheory/Ordinal/NaturalOps.lean b/Mathlib/SetTheory/Ordinal/NaturalOps.lean index 7463af175343e..58624fce3613b 100644 --- a/Mathlib/SetTheory/Ordinal/NaturalOps.lean +++ b/Mathlib/SetTheory/Ordinal/NaturalOps.lean @@ -78,10 +78,9 @@ open Ordinal theorem toOrdinal_symm_eq : NatOrdinal.toOrdinal.symm = Ordinal.toNatOrdinal := rfl --- Porting note: used to use dot notation, but doesn't work in Lean 4 with `OrderIso` @[simp] -theorem toOrdinal_toNatOrdinal (a : NatOrdinal) : - Ordinal.toNatOrdinal (NatOrdinal.toOrdinal a) = a := rfl +theorem toOrdinal_toNatOrdinal (a : NatOrdinal) : a.toOrdinal.toNatOrdinal = a := + rfl theorem lt_wf : @WellFounded NatOrdinal (· < ·) := Ordinal.lt_wf @@ -143,7 +142,7 @@ theorem toNatOrdinal_symm_eq : toNatOrdinal.symm = NatOrdinal.toOrdinal := rfl @[simp] -theorem toNatOrdinal_toOrdinal (a : Ordinal) : NatOrdinal.toOrdinal (toNatOrdinal a) = a := +theorem toNatOrdinal_toOrdinal (a : Ordinal) : a.toNatOrdinal.toOrdinal = a := rfl @[simp] @@ -440,8 +439,10 @@ theorem nadd_right_comm : ∀ a b c, a ♯ b ♯ c = a ♯ c ♯ b := variable {a b c d : Ordinal.{u}} +@[deprecated "avoid using the definition of `nmul` directly" (since := "2024-11-19")] theorem nmul_def (a b : Ordinal) : - a ⨳ b = sInf {c | ∀ a' < a, ∀ b' < b, a' ⨳ b ♯ a ⨳ b' < c ♯ a' ⨳ b'} := by rw [nmul] + a ⨳ b = sInf {c | ∀ a' < a, ∀ b' < b, a' ⨳ b ♯ a ⨳ b' < c ♯ a' ⨳ b'} := by + rw [nmul] /-- The set in the definition of `nmul` is nonempty. -/ private theorem nmul_nonempty (a b : Ordinal.{u}) : @@ -453,7 +454,7 @@ private theorem nmul_nonempty (a b : Ordinal.{u}) : theorem nmul_nadd_lt {a' b' : Ordinal} (ha : a' < a) (hb : b' < b) : a' ⨳ b ♯ a ⨳ b' < a ⨳ b ♯ a' ⨳ b' := by - rw [nmul_def a b] + conv_rhs => rw [nmul] exact csInf_mem (nmul_nonempty a b) a' ha b' hb theorem nmul_nadd_le {a' b' : Ordinal} (ha : a' ≤ a) (hb : b' ≤ b) : @@ -478,11 +479,9 @@ theorem nmul_le_iff : a ⨳ b ≤ c ↔ ∀ a' < a, ∀ b' < b, a' ⨳ b ♯ a theorem nmul_comm (a b) : a ⨳ b = b ⨳ a := by rw [nmul, nmul] congr; ext x; constructor <;> intro H c hc d hd - -- Porting note: had to add additional arguments to `nmul_comm` here - -- for the termination checker. - · rw [nadd_comm, ← nmul_comm d b, ← nmul_comm a c, ← nmul_comm d] + · rw [nadd_comm, ← nmul_comm, ← nmul_comm a, ← nmul_comm d] exact H _ hd _ hc - · rw [nadd_comm, nmul_comm a d, nmul_comm c, nmul_comm c] + · rw [nadd_comm, nmul_comm, nmul_comm c, nmul_comm c] exact H _ hd _ hc termination_by (a, b) @@ -497,17 +496,13 @@ theorem zero_nmul (a) : 0 ⨳ a = 0 := by rw [nmul_comm, nmul_zero] @[simp] theorem nmul_one (a : Ordinal) : a ⨳ 1 = a := by rw [nmul] - simp only [lt_one_iff_zero, forall_eq, nmul_zero, nadd_zero] - convert csInf_Ici (α := Ordinal) + convert csInf_Ici ext b - -- Porting note: added this `simp` line, as the result from `convert` - -- is slightly different. - simp only [Set.mem_setOf_eq, Set.mem_Ici] - refine ⟨fun H => le_of_forall_lt fun c hc => ?_, fun ha c hc => ?_⟩ + refine ⟨fun H ↦ le_of_forall_lt (a := a) fun c hc ↦ ?_, fun ha c hc ↦ ?_⟩ -- Porting note: had to add arguments to `nmul_one` in the next two lines -- for the termination checker. - · simpa only [nmul_one c] using H c hc - · simpa only [nmul_one c] using hc.trans_le ha + · simpa [nmul_one c] using H c hc + · simpa [nmul_one c] using hc.trans_le ha termination_by a @[simp] @@ -537,19 +532,16 @@ alias nmul_le_nmul_of_nonneg_right := nmul_le_nmul_right theorem nmul_nadd (a b c : Ordinal) : a ⨳ (b ♯ c) = a ⨳ b ♯ a ⨳ c := by refine le_antisymm (nmul_le_iff.2 fun a' ha d hd => ?_) (nadd_le_iff.2 ⟨fun d hd => ?_, fun d hd => ?_⟩) - · -- Porting note: adding arguments to `nmul_nadd` for the termination checker. - rw [nmul_nadd a' b c] + · rw [nmul_nadd] rcases lt_nadd_iff.1 hd with (⟨b', hb, hd⟩ | ⟨c', hc, hd⟩) · have := nadd_lt_nadd_of_lt_of_le (nmul_nadd_lt ha hb) (nmul_nadd_le ha.le hd) - -- Porting note: adding arguments to `nmul_nadd` for the termination checker. - rw [nmul_nadd a' b' c, nmul_nadd a b' c] at this + rw [nmul_nadd, nmul_nadd] at this simp only [nadd_assoc] at this rwa [nadd_left_comm, nadd_left_comm _ (a ⨳ b'), nadd_left_comm (a ⨳ b), nadd_lt_nadd_iff_left, nadd_left_comm (a' ⨳ b), nadd_left_comm (a ⨳ b), nadd_lt_nadd_iff_left, ← nadd_assoc, ← nadd_assoc] at this · have := nadd_lt_nadd_of_le_of_lt (nmul_nadd_le ha.le hd) (nmul_nadd_lt ha hc) - -- Porting note: adding arguments to `nmul_nadd` for the termination checker. - rw [nmul_nadd a' b c', nmul_nadd a b c'] at this + rw [nmul_nadd, nmul_nadd] at this simp only [nadd_assoc] at this rwa [nadd_left_comm, nadd_comm (a ⨳ c), nadd_left_comm (a' ⨳ d), nadd_left_comm (a ⨳ c'), nadd_left_comm (a ⨳ b), nadd_lt_nadd_iff_left, nadd_comm (a' ⨳ c), nadd_left_comm (a ⨳ d), @@ -557,8 +549,7 @@ theorem nmul_nadd (a b c : Ordinal) : a ⨳ (b ♯ c) = a ⨳ b ♯ a ⨳ c := b nadd_comm (a' ⨳ d), ← nadd_assoc, ← nadd_assoc] at this · rcases lt_nmul_iff.1 hd with ⟨a', ha, b', hb, hd⟩ have := nadd_lt_nadd_of_le_of_lt hd (nmul_nadd_lt ha (nadd_lt_nadd_right hb c)) - -- Porting note: adding arguments to `nmul_nadd` for the termination checker. - rw [nmul_nadd a' b c, nmul_nadd a b' c, nmul_nadd a'] at this + rw [nmul_nadd, nmul_nadd, nmul_nadd a'] at this simp only [nadd_assoc] at this rwa [nadd_left_comm (a' ⨳ b'), nadd_left_comm, nadd_lt_nadd_iff_left, nadd_left_comm, nadd_left_comm _ (a' ⨳ b'), nadd_left_comm (a ⨳ b'), nadd_lt_nadd_iff_left, @@ -566,8 +557,7 @@ theorem nmul_nadd (a b c : Ordinal) : a ⨳ (b ♯ c) = a ⨳ b ♯ a ⨳ c := b nadd_comm _ (a' ⨳ c), nadd_lt_nadd_iff_left] at this · rcases lt_nmul_iff.1 hd with ⟨a', ha, c', hc, hd⟩ have := nadd_lt_nadd_of_lt_of_le (nmul_nadd_lt ha (nadd_lt_nadd_left hc b)) hd - -- Porting note: adding arguments to `nmul_nadd` for the termination checker. - rw [nmul_nadd a' b c, nmul_nadd a b c', nmul_nadd a'] at this + rw [nmul_nadd, nmul_nadd, nmul_nadd a'] at this simp only [nadd_assoc] at this rwa [nadd_left_comm _ (a' ⨳ b), nadd_lt_nadd_iff_left, nadd_left_comm (a' ⨳ c'), nadd_left_comm _ (a' ⨳ c), nadd_lt_nadd_iff_left, nadd_left_comm, nadd_comm (a' ⨳ c'), @@ -588,26 +578,25 @@ theorem nmul_nadd_le₃ {a' b' c' : Ordinal} (ha : a' ≤ a) (hb : b' ≤ b) (hc a ⨳ b ⨳ c ♯ a' ⨳ b' ⨳ c ♯ a' ⨳ b ⨳ c' ♯ a ⨳ b' ⨳ c' := by simpa only [nadd_nmul, ← nadd_assoc] using nmul_nadd_le (nmul_nadd_le ha hb) hc -theorem nmul_nadd_lt₃' {a' b' c' : Ordinal} (ha : a' < a) (hb : b' < b) (hc : c' < c) : +private theorem nmul_nadd_lt₃' {a' b' c' : Ordinal} (ha : a' < a) (hb : b' < b) (hc : c' < c) : a' ⨳ (b ⨳ c) ♯ a ⨳ (b' ⨳ c) ♯ a ⨳ (b ⨳ c') ♯ a' ⨳ (b' ⨳ c') < a ⨳ (b ⨳ c) ♯ a' ⨳ (b' ⨳ c) ♯ a' ⨳ (b ⨳ c') ♯ a ⨳ (b' ⨳ c') := by simp only [nmul_comm _ (_ ⨳ _)] convert nmul_nadd_lt₃ hb hc ha using 1 <;> - · simp only [nadd_eq_add, NatOrdinal.toOrdinal_toNatOrdinal]; abel_nf + (simp only [nadd_eq_add, NatOrdinal.toOrdinal_toNatOrdinal]; abel_nf) +@[deprecated nmul_nadd_le₃ (since := "2024-11-19")] theorem nmul_nadd_le₃' {a' b' c' : Ordinal} (ha : a' ≤ a) (hb : b' ≤ b) (hc : c' ≤ c) : a' ⨳ (b ⨳ c) ♯ a ⨳ (b' ⨳ c) ♯ a ⨳ (b ⨳ c') ♯ a' ⨳ (b' ⨳ c') ≤ a ⨳ (b ⨳ c) ♯ a' ⨳ (b' ⨳ c) ♯ a' ⨳ (b ⨳ c') ♯ a ⨳ (b' ⨳ c') := by simp only [nmul_comm _ (_ ⨳ _)] convert nmul_nadd_le₃ hb hc ha using 1 <;> - · simp only [nadd_eq_add, NatOrdinal.toOrdinal_toNatOrdinal]; abel_nf + (simp only [nadd_eq_add, NatOrdinal.toOrdinal_toNatOrdinal]; abel_nf) -theorem lt_nmul_iff₃ : - d < a ⨳ b ⨳ c ↔ - ∃ a' < a, ∃ b' < b, ∃ c' < c, - d ♯ a' ⨳ b' ⨳ c ♯ a' ⨳ b ⨳ c' ♯ a ⨳ b' ⨳ c' ≤ - a' ⨳ b ⨳ c ♯ a ⨳ b' ⨳ c ♯ a ⨳ b ⨳ c' ♯ a' ⨳ b' ⨳ c' := by - refine ⟨fun h => ?_, ?_⟩ +theorem lt_nmul_iff₃ : d < a ⨳ b ⨳ c ↔ ∃ a' < a, ∃ b' < b, ∃ c' < c, + d ♯ a' ⨳ b' ⨳ c ♯ a' ⨳ b ⨳ c' ♯ a ⨳ b' ⨳ c' ≤ + a' ⨳ b ⨳ c ♯ a ⨳ b' ⨳ c ♯ a ⨳ b ⨳ c' ♯ a' ⨳ b' ⨳ c' := by + refine ⟨fun h ↦ ?_, fun ⟨a', ha, b', hb, c', hc, h⟩ ↦ ?_⟩ · rcases lt_nmul_iff.1 h with ⟨e, he, c', hc, H₁⟩ rcases lt_nmul_iff.1 he with ⟨a', ha, b', hb, H₂⟩ refine ⟨a', ha, b', hb, c', hc, ?_⟩ @@ -617,52 +606,38 @@ theorem lt_nmul_iff₃ : nadd_left_comm (a ⨳ b' ⨳ c), nadd_left_comm (a' ⨳ b ⨳ c), nadd_left_comm (a ⨳ b ⨳ c'), nadd_le_nadd_iff_left, nadd_left_comm (a ⨳ b ⨳ c'), nadd_left_comm (a ⨳ b ⨳ c')] at this simpa only [nadd_assoc] - · rintro ⟨a', ha, b', hb, c', hc, h⟩ - have := h.trans_lt (nmul_nadd_lt₃ ha hb hc) - repeat' rw [nadd_lt_nadd_iff_right] at this + · have := h.trans_lt (nmul_nadd_lt₃ ha hb hc) + repeat rw [nadd_lt_nadd_iff_right] at this assumption -theorem nmul_le_iff₃ : - a ⨳ b ⨳ c ≤ d ↔ - ∀ a' < a, ∀ b' < b, ∀ c' < c, - a' ⨳ b ⨳ c ♯ a ⨳ b' ⨳ c ♯ a ⨳ b ⨳ c' ♯ a' ⨳ b' ⨳ c' < - d ♯ a' ⨳ b' ⨳ c ♯ a' ⨳ b ⨳ c' ♯ a ⨳ b' ⨳ c' := by - rw [← not_iff_not]; simp [lt_nmul_iff₃] - -theorem lt_nmul_iff₃' : - d < a ⨳ (b ⨳ c) ↔ - ∃ a' < a, ∃ b' < b, ∃ c' < c, - d ♯ a' ⨳ (b' ⨳ c) ♯ a' ⨳ (b ⨳ c') ♯ a ⨳ (b' ⨳ c') ≤ - a' ⨳ (b ⨳ c) ♯ a ⨳ (b' ⨳ c) ♯ a ⨳ (b ⨳ c') ♯ a' ⨳ (b' ⨳ c') := by - simp only [nmul_comm _ (_ ⨳ _), lt_nmul_iff₃, nadd_eq_add, NatOrdinal.toOrdinal_toNatOrdinal] - constructor <;> rintro ⟨b', hb, c', hc, a', ha, h⟩ - · use a', ha, b', hb, c', hc; convert h using 1 <;> abel_nf - · use c', hc, a', ha, b', hb; convert h using 1 <;> abel_nf - -theorem nmul_le_iff₃' : - a ⨳ (b ⨳ c) ≤ d ↔ - ∀ a' < a, ∀ b' < b, ∀ c' < c, - a' ⨳ (b ⨳ c) ♯ a ⨳ (b' ⨳ c) ♯ a ⨳ (b ⨳ c') ♯ a' ⨳ (b' ⨳ c') < - d ♯ a' ⨳ (b' ⨳ c) ♯ a' ⨳ (b ⨳ c') ♯ a ⨳ (b' ⨳ c') := by - rw [← not_iff_not]; simp [lt_nmul_iff₃'] +theorem nmul_le_iff₃ : a ⨳ b ⨳ c ≤ d ↔ ∀ a' < a, ∀ b' < b, ∀ c' < c, + a' ⨳ b ⨳ c ♯ a ⨳ b' ⨳ c ♯ a ⨳ b ⨳ c' ♯ a' ⨳ b' ⨳ c' < + d ♯ a' ⨳ b' ⨳ c ♯ a' ⨳ b ⨳ c' ♯ a ⨳ b' ⨳ c' := by + simpa using lt_nmul_iff₃.not + +private theorem nmul_le_iff₃' : a ⨳ (b ⨳ c) ≤ d ↔ ∀ a' < a, ∀ b' < b, ∀ c' < c, + a' ⨳ (b ⨳ c) ♯ a ⨳ (b' ⨳ c) ♯ a ⨳ (b ⨳ c') ♯ a' ⨳ (b' ⨳ c') < + d ♯ a' ⨳ (b' ⨳ c) ♯ a' ⨳ (b ⨳ c') ♯ a ⨳ (b' ⨳ c') := by + simp only [nmul_comm _ (_ ⨳ _), nmul_le_iff₃, nadd_eq_add, toOrdinal_toNatOrdinal] + constructor <;> intro h a' ha b' hb c' hc + · convert h b' hb c' hc a' ha using 1 <;> abel_nf + · convert h c' hc a' ha b' hb using 1 <;> abel_nf + +@[deprecated lt_nmul_iff₃ (since := "2024-11-19")] +theorem lt_nmul_iff₃' : d < a ⨳ (b ⨳ c) ↔ ∃ a' < a, ∃ b' < b, ∃ c' < c, + d ♯ a' ⨳ (b' ⨳ c) ♯ a' ⨳ (b ⨳ c') ♯ a ⨳ (b' ⨳ c') ≤ + a' ⨳ (b ⨳ c) ♯ a ⨳ (b' ⨳ c) ♯ a ⨳ (b ⨳ c') ♯ a' ⨳ (b' ⨳ c') := by + simpa using nmul_le_iff₃'.not theorem nmul_assoc (a b c : Ordinal) : a ⨳ b ⨳ c = a ⨳ (b ⨳ c) := by apply le_antisymm · rw [nmul_le_iff₃] intro a' ha b' hb c' hc - -- Porting note: the next line was just - -- repeat' rw [nmul_assoc] - -- but we need to spell out the arguments for the termination checker. - rw [nmul_assoc a' b c, nmul_assoc a b' c, nmul_assoc a b c', nmul_assoc a' b' c', - nmul_assoc a' b' c, nmul_assoc a' b c', nmul_assoc a b' c'] + repeat rw [nmul_assoc] exact nmul_nadd_lt₃' ha hb hc · rw [nmul_le_iff₃'] intro a' ha b' hb c' hc - -- Porting note: the next line was just - -- repeat' rw [← nmul_assoc] - -- but we need to spell out the arguments for the termination checker. - rw [← nmul_assoc a' b c, ← nmul_assoc a b' c, ← nmul_assoc a b c', ← nmul_assoc a' b' c', - ← nmul_assoc a' b' c, ← nmul_assoc a' b c', ← nmul_assoc a b' c'] + repeat rw [← nmul_assoc] exact nmul_nadd_lt₃ ha hb hc termination_by (a, b, c) From 9e1076898ba7f097186b98ffe8a59a2ef8b61200 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 16 Dec 2024 00:26:38 +0000 Subject: [PATCH 728/829] feat(SetTheory/Cardinal/Arithmetic): cardinality of ordinal exponential (#17813) We prove variations on `(a ^ b).card = max a.card b.card`. --- Mathlib/SetTheory/Cardinal/Aleph.lean | 3 + Mathlib/SetTheory/Cardinal/Arithmetic.lean | 78 ++++++++++++++++++++++ Mathlib/SetTheory/Ordinal/Arithmetic.lean | 5 ++ Mathlib/SetTheory/Ordinal/Principal.lean | 3 - 4 files changed, 86 insertions(+), 3 deletions(-) diff --git a/Mathlib/SetTheory/Cardinal/Aleph.lean b/Mathlib/SetTheory/Cardinal/Aleph.lean index db72a7115f40d..896cc181244ef 100644 --- a/Mathlib/SetTheory/Cardinal/Aleph.lean +++ b/Mathlib/SetTheory/Cardinal/Aleph.lean @@ -498,6 +498,9 @@ theorem aleph1_eq_lift {c : Cardinal.{u}} : ℵ₁ = lift.{v} c ↔ ℵ₁ = c : theorem lift_eq_aleph1 {c : Cardinal.{u}} : lift.{v} c = ℵ₁ ↔ c = ℵ₁ := by simpa using lift_inj (b := ℵ₁) +theorem lt_omega_iff_card_lt {x o : Ordinal} : x < ω_ o ↔ x.card < ℵ_ o := by + rw [← (isInitial_omega o).card_lt_card, card_omega] + section deprecated set_option linter.docPrime false diff --git a/Mathlib/SetTheory/Cardinal/Arithmetic.lean b/Mathlib/SetTheory/Cardinal/Arithmetic.lean index 8f3b807c4c581..88177355e75c5 100644 --- a/Mathlib/SetTheory/Cardinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Cardinal/Arithmetic.lean @@ -881,4 +881,82 @@ theorem card_iSup_Iio_le_card_mul_iSup {o : Ordinal.{u}} (f : Iio o → Ordinal. · exact mk_toType o · exact (enumIsoToType o).symm.iSup_comp (g := fun x ↦ (f x).card) +theorem card_opow_le_of_omega0_le_left {a : Ordinal} (ha : ω ≤ a) (b : Ordinal) : + (a ^ b).card ≤ max a.card b.card := by + refine limitRecOn b ?_ ?_ ?_ + · simpa using one_lt_omega0.le.trans ha + · intro b IH + rw [opow_succ, card_mul, card_succ, Cardinal.mul_eq_max_of_aleph0_le_right, max_comm] + · apply (max_le_max_left _ IH).trans + rw [← max_assoc, max_self] + exact max_le_max_left _ le_self_add + · rw [ne_eq, card_eq_zero, opow_eq_zero] + rintro ⟨rfl, -⟩ + cases omega0_pos.not_le ha + · rwa [aleph0_le_card] + · intro b hb IH + rw [(isNormal_opow (one_lt_omega0.trans_le ha)).apply_of_isLimit hb] + apply (card_iSup_Iio_le_card_mul_iSup _).trans + rw [Cardinal.lift_id, Cardinal.mul_eq_max_of_aleph0_le_right, max_comm] + · apply max_le _ (le_max_right _ _) + apply ciSup_le' + intro c + exact (IH c.1 c.2).trans (max_le_max_left _ (card_le_card c.2.le)) + · simpa using hb.pos.ne' + · refine le_ciSup_of_le ?_ ⟨1, one_lt_omega0.trans_le <| omega0_le_of_isLimit hb⟩ ?_ + · exact Cardinal.bddAbove_of_small _ + · simpa + +theorem card_opow_le_of_omega0_le_right (a : Ordinal) {b : Ordinal} (hb : ω ≤ b) : + (a ^ b).card ≤ max a.card b.card := by + obtain ⟨n, rfl⟩ | ha := eq_nat_or_omega0_le a + · apply (card_le_card <| opow_le_opow_left b (nat_lt_omega0 n).le).trans + apply (card_opow_le_of_omega0_le_left le_rfl _).trans + simp [hb] + · exact card_opow_le_of_omega0_le_left ha b + +theorem card_opow_le (a b : Ordinal) : (a ^ b).card ≤ max ℵ₀ (max a.card b.card) := by + obtain ⟨n, rfl⟩ | ha := eq_nat_or_omega0_le a + · obtain ⟨m, rfl⟩ | hb := eq_nat_or_omega0_le b + · rw [← natCast_opow, card_nat] + exact le_max_of_le_left (nat_lt_aleph0 _).le + · exact (card_opow_le_of_omega0_le_right _ hb).trans (le_max_right _ _) + · exact (card_opow_le_of_omega0_le_left ha _).trans (le_max_right _ _) + +theorem card_opow_eq_of_omega0_le_left {a b : Ordinal} (ha : ω ≤ a) (hb : 0 < b) : + (a ^ b).card = max a.card b.card := by + apply (card_opow_le_of_omega0_le_left ha b).antisymm (max_le _ _) <;> apply card_le_card + · exact left_le_opow a hb + · exact right_le_opow b (one_lt_omega0.trans_le ha) + +theorem card_opow_eq_of_omega0_le_right {a b : Ordinal} (ha : 1 < a) (hb : ω ≤ b) : + (a ^ b).card = max a.card b.card := by + apply (card_opow_le_of_omega0_le_right a hb).antisymm (max_le _ _) <;> apply card_le_card + · exact left_le_opow a (omega0_pos.trans_le hb) + · exact right_le_opow b ha + +theorem card_omega0_opow {a : Ordinal} (h : a ≠ 0) : card (ω ^ a) = max ℵ₀ a.card := by + rw [card_opow_eq_of_omega0_le_left le_rfl h.bot_lt, card_omega0] + +theorem card_opow_omega0 {a : Ordinal} (h : 1 < a) : card (a ^ ω) = max ℵ₀ a.card := by + rw [card_opow_eq_of_omega0_le_right h le_rfl, card_omega0, max_comm] + +theorem principal_opow_omega (o : Ordinal) : Principal (· ^ ·) (ω_ o) := by + obtain rfl | ho := Ordinal.eq_zero_or_pos o + · rw [omega_zero] + exact principal_opow_omega0 + · intro a b ha hb + rw [lt_omega_iff_card_lt] at ha hb ⊢ + apply (card_opow_le a b).trans_lt (max_lt _ (max_lt ha hb)) + rwa [← aleph_zero, aleph_lt_aleph] + +theorem IsInitial.principal_opow {o : Ordinal} (h : IsInitial o) (ho : ω ≤ o) : + Principal (· ^ ·) o := by + obtain ⟨a, rfl⟩ := mem_range_omega_iff.2 ⟨ho, h⟩ + exact principal_opow_omega a + +theorem principal_opow_ord {c : Cardinal} (hc : ℵ₀ ≤ c) : Principal (· ^ ·) c.ord := by + apply (isInitial_ord c).principal_opow + rwa [omega0_le_ord] + end Ordinal diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index bd4d0884f8560..2b736379c5094 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -2351,6 +2351,11 @@ theorem nat_lt_omega0 (n : ℕ) : ↑n < ω := @[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias nat_lt_omega := nat_lt_omega0 +theorem eq_nat_or_omega0_le (o : Ordinal) : (∃ n : ℕ, o = n) ∨ ω ≤ o := by + obtain ho | ho := lt_or_le o ω + · exact Or.inl <| lt_omega0.1 ho + · exact Or.inr ho + theorem omega0_pos : 0 < ω := nat_lt_omega0 0 diff --git a/Mathlib/SetTheory/Ordinal/Principal.lean b/Mathlib/SetTheory/Ordinal/Principal.lean index d993ba3d10462..1500096b7b3a1 100644 --- a/Mathlib/SetTheory/Ordinal/Principal.lean +++ b/Mathlib/SetTheory/Ordinal/Principal.lean @@ -459,9 +459,6 @@ theorem principal_opow_omega0 : Principal (· ^ ·) ω := fun a b ha hb => simp_rw [← natCast_opow] apply nat_lt_omega0 -@[deprecated (since := "2024-09-30")] -alias principal_opow_omega := principal_opow_omega0 - theorem opow_omega0 (a1 : 1 < a) (h : a < ω) : a ^ ω = ω := ((opow_le_of_limit (one_le_iff_ne_zero.1 <| le_of_lt a1) isLimit_omega0).2 fun _ hb => (principal_opow_omega0 h hb).le).antisymm From dd3f45bad90e99f476e1c040bbbdb2fa67c63261 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 16 Dec 2024 00:47:04 +0000 Subject: [PATCH 729/829] =?UTF-8?q?feat(SetTheory/Ordinal/Arithmetic):=20`?= =?UTF-8?q?(sInf=20s=E1=B6=9C).card=20=E2=89=A4=20#s`=20(#18780)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The cardinality of the least ordinal not in a set is at most the cardinality of the set. --- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 29 ++++++++++++++++------- 1 file changed, 20 insertions(+), 9 deletions(-) diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 2b736379c5094..7fd96b5c22941 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -1466,17 +1466,28 @@ theorem iSup_ord {ι} {f : ι → Cardinal} (hf : BddAbove (range f)) : conv_lhs => change range (ord ∘ f) rw [range_comp] +theorem lift_card_sInf_compl_le (s : Set Ordinal.{u}) : + Cardinal.lift.{u + 1} (sInf sᶜ).card ≤ #s := by + rw [← mk_Iio_ordinal] + refine mk_le_mk_of_subset fun x (hx : x < _) ↦ ?_ + rw [← not_not_mem] + exact not_mem_of_lt_csInf' hx + +theorem card_sInf_range_compl_le_lift {ι : Type u} (f : ι → Ordinal.{max u v}) : + (sInf (range f)ᶜ).card ≤ Cardinal.lift.{v} #ι := by + rw [← Cardinal.lift_le.{max u v + 1}, Cardinal.lift_lift] + apply (lift_card_sInf_compl_le _).trans + rw [← Cardinal.lift_id'.{u, max u v + 1} #(range _)] + exact mk_range_le_lift + +theorem card_sInf_range_compl_le {ι : Type u} (f : ι → Ordinal.{u}) : + (sInf (range f)ᶜ).card ≤ #ι := + Cardinal.lift_id #ι ▸ card_sInf_range_compl_le_lift f + theorem sInf_compl_lt_lift_ord_succ {ι : Type u} (f : ι → Ordinal.{max u v}) : sInf (range f)ᶜ < lift.{v} (succ #ι).ord := by - by_contra! h - have : Iio (lift.{v} (succ #ι).ord) ⊆ range f := by - intro o ho - have := not_mem_of_lt_csInf' (ho.trans_le h) - rwa [not_mem_compl_iff] at this - have := mk_le_mk_of_subset this - rw [mk_Iio_ordinal, ← lift_card, Cardinal.lift_lift, card_ord, Cardinal.lift_succ, - succ_le_iff, ← Cardinal.lift_id'.{u, max (u + 1) (v + 1)} #_] at this - exact this.not_le mk_range_le_lift + rw [lift_ord, Cardinal.lift_succ, ← card_le_iff] + exact card_sInf_range_compl_le_lift f theorem sInf_compl_lt_ord_succ {ι : Type u} (f : ι → Ordinal.{u}) : sInf (range f)ᶜ < (succ #ι).ord := From d3319caa12aeaa10c9b9e85d55c40af16dbc52a7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Iv=C3=A1n=20Renison?= <85908989+IvanRenison@users.noreply.github.com> Date: Mon, 16 Dec 2024 01:07:35 +0000 Subject: [PATCH 730/829] chore(Algebra/Group/Nat): Remove some unnecessary lemmas arguments (#19981) --- Mathlib/Algebra/Group/Nat/Even.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/Algebra/Group/Nat/Even.lean b/Mathlib/Algebra/Group/Nat/Even.lean index 3377025e9fb61..4dd418f24a6cb 100644 --- a/Mathlib/Algebra/Group/Nat/Even.lean +++ b/Mathlib/Algebra/Group/Nat/Even.lean @@ -44,10 +44,10 @@ lemma not_even_iff : ¬ Even n ↔ n % 2 = 1 := by rw [even_iff, mod_two_not_eq_ @[parity_simps] lemma even_add_one : Even (n + 1) ↔ ¬Even n := by simp [even_add] -lemma succ_mod_two_eq_zero_iff {m : ℕ} : (m + 1) % 2 = 0 ↔ m % 2 = 1 := by +lemma succ_mod_two_eq_zero_iff : (m + 1) % 2 = 0 ↔ m % 2 = 1 := by simp [← Nat.even_iff, ← Nat.not_even_iff, parity_simps] -lemma succ_mod_two_eq_one_iff {m : ℕ} : (m + 1) % 2 = 1 ↔ m % 2 = 0 := by +lemma succ_mod_two_eq_one_iff : (m + 1) % 2 = 1 ↔ m % 2 = 0 := by simp [← Nat.even_iff, ← Nat.not_even_iff, parity_simps] lemma two_not_dvd_two_mul_add_one (n : ℕ) : ¬2 ∣ 2 * n + 1 := by simp [add_mod] @@ -82,13 +82,13 @@ lemma two_mul_div_two_of_even : Even n → 2 * (n / 2) = n := fun h ↦ lemma div_two_mul_two_of_even : Even n → n / 2 * 2 = n := fun h ↦ Nat.div_mul_cancel ((even_iff_exists_two_nsmul _).1 h) -theorem one_lt_of_ne_zero_of_even {n : ℕ} (h0 : n ≠ 0) (hn : Even n) : 1 < n := by +theorem one_lt_of_ne_zero_of_even (h0 : n ≠ 0) (hn : Even n) : 1 < n := by refine Nat.one_lt_iff_ne_zero_and_ne_one.mpr (And.intro h0 ?_) intro h rw [h] at hn exact Nat.not_even_one hn -theorem add_one_lt_of_even {n m : ℕ} (hn : Even n) (hm : Even m) (hnm : n < m) : +theorem add_one_lt_of_even (hn : Even n) (hm : Even m) (hnm : n < m) : n + 1 < m := by rcases hn with ⟨n, rfl⟩ rcases hm with ⟨m, rfl⟩ From 3f813de52d8cffaa73e27edd62364eec90eac633 Mon Sep 17 00:00:00 2001 From: "Tristan F.-R." Date: Mon, 16 Dec 2024 01:30:05 +0000 Subject: [PATCH 731/829] feat(SetTheory/Game/PGame): identical_of_eq (#19983) To complement [`equiv_of_eq`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/SetTheory/Game/PGame.html#SetTheory.PGame.equiv_of_eq). Co-authored-by: Tristan F. --- Mathlib/SetTheory/Game/PGame.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/Mathlib/SetTheory/Game/PGame.lean b/Mathlib/SetTheory/Game/PGame.lean index d1e221f7c4878..91a9c5e7af475 100644 --- a/Mathlib/SetTheory/Game/PGame.lean +++ b/Mathlib/SetTheory/Game/PGame.lean @@ -414,6 +414,8 @@ lemma Identical.moveRight : ∀ {x y}, x ≡ y → ∀ i, ∃ j, x.moveRight i ≡ y.moveRight j | mk _ _ _ _, mk _ _ _ _, ⟨_, hr⟩, i => hr.1 i +theorem identical_of_eq {x y : PGame} (h : x = y) : x ≡ y := by subst h; rfl + /-- Uses `∈ₗ` and `∈ᵣ` instead of `≡`. -/ theorem identical_iff' : ∀ {x y : PGame}, x ≡ y ↔ ((∀ i, x.moveLeft i ∈ₗ y) ∧ (∀ j, y.moveLeft j ∈ₗ x)) ∧ From aee9be2beb5b2fc03a8d7c6a64d4683f8c365f2a Mon Sep 17 00:00:00 2001 From: smorel394 Date: Mon, 16 Dec 2024 08:38:43 +0000 Subject: [PATCH 732/829] feat(Algebra/Category/MonCat/Limits): the forgetful functor on the category of monoids creates all limits (#19965) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The forgetful functor on the category of monoids (commutative or not) creates all limits. The only non-trivial lemma is `MonCat.forget_createsLimit`, but even that is easy because it uses that limits in the category of monoids are literally constructed by taking limits in the category of types and putting a monoid structure on the result. The only "innovation" is noticing that the limit of `F` existing in the category of types already forces the sections of `F ⋙ forget MonCat``` to be small, so we don't need to impose any smallness condition on our functors. - [x] depends on: #19964 [the forgetful functor on `MonCat` is corepresentable] --- Mathlib/Algebra/Category/MonCat/Limits.lean | 69 ++++++++++++++++++++- 1 file changed, 67 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Category/MonCat/Limits.lean b/Mathlib/Algebra/Category/MonCat/Limits.lean index 8145dc0c3be8e..ef95015ba092d 100644 --- a/Mathlib/Algebra/Category/MonCat/Limits.lean +++ b/Mathlib/Algebra/Category/MonCat/Limits.lean @@ -4,8 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ import Mathlib.Algebra.Category.MonCat.Basic -import Mathlib.Algebra.Group.Pi.Lemmas -import Mathlib.Algebra.Group.Submonoid.Operations import Mathlib.CategoryTheory.Limits.Creates import Mathlib.CategoryTheory.Limits.Types import Mathlib.Logic.Small.Group @@ -157,6 +155,48 @@ noncomputable instance forget_preservesLimitsOfSize [UnivLE.{v, u}] : noncomputable instance forget_preservesLimits : PreservesLimits (forget MonCat.{u}) := MonCat.forget_preservesLimitsOfSize.{u, u} +@[to_additive] +noncomputable instance forget_createsLimit : + CreatesLimit F (forget MonCat.{u}) := by + apply createsLimitOfReflectsIso + intro c t + have : Small.{u} (Functor.sections (F ⋙ forget MonCat)) := + (Types.hasLimit_iff_small_sections _).mp (HasLimit.mk {cone := c, isLimit := t}) + refine LiftsToLimit.mk (LiftableCone.mk + {pt := MonCat.of (Types.Small.limitCone (F ⋙ forget MonCat)).pt, π := NatTrans.mk + (limitπMonoidHom F) (MonCat.HasLimits.limitCone F).π.naturality} (Cones.ext + ((Types.isLimitEquivSections t).trans (equivShrink _)).symm.toIso + (fun _ ↦ funext (fun _ ↦ by simp; rfl)))) ?_ + refine IsLimit.ofFaithful (forget MonCat.{u}) (Types.Small.limitConeIsLimit.{v,u} _) ?_ ?_ + · intro _ + refine {toFun := (Types.Small.limitConeIsLimit.{v,u} _).lift ((forget MonCat).mapCone _), + map_one' := by simp; rfl, map_mul' := ?_ } + · intro x y + simp only [Types.Small.limitConeIsLimit_lift, Functor.comp_obj, Functor.mapCone_pt, + Functor.mapCone_π_app, forget_map, map_mul, mul_of] + congr + simp only [Functor.comp_obj, Equiv.symm_apply_apply] + rfl + · exact fun _ ↦ rfl + +@[to_additive] +noncomputable instance forget_createsLimitsOfShape : + CreatesLimitsOfShape J (forget MonCat.{u}) where + CreatesLimit := inferInstance + +/-- The forgetful functor from monoids to types preserves all limits. +-/ +@[to_additive +"The forgetful functor from additive monoids to types preserves all limits." +] +noncomputable instance forget_createsLimitsOfSize : + CreatesLimitsOfSize.{w,v} (forget MonCat.{u}) where + CreatesLimitsOfShape := inferInstance + +@[to_additive] +noncomputable instance forget_createsLimits : + CreatesLimits (forget MonCat.{u}) := MonCat.forget_createsLimitsOfSize.{u,u} + end MonCat open MonCat @@ -292,4 +332,29 @@ instance _root_.AddCommMonCat.forget_preservesLimits : instance forget_preservesLimits : PreservesLimits (forget CommMonCat.{u}) := CommMonCat.forget_preservesLimitsOfSize.{u, u} +@[to_additive] +noncomputable instance forget_createsLimit : + CreatesLimit F (forget CommMonCat.{u}) := by + set e : forget CommMonCat.{u} ≅ forget₂ CommMonCat.{u} MonCat.{u} ⋙ forget MonCat.{u} := + NatIso.ofComponents (fun _ ↦ Iso.refl _) (fun _ ↦ rfl) + exact createsLimitOfNatIso e.symm + +@[to_additive] +noncomputable instance forget_createsLimitsOfShape : + CreatesLimitsOfShape J (forget MonCat.{u}) where + CreatesLimit := inferInstance + +/-- The forgetful functor from commutative monoids to types preserves all limits. +-/ +@[to_additive +"The forgetful functor from commutative additive monoids to types preserves all limits." +] +noncomputable instance forget_createsLimitsOfSize : + CreatesLimitsOfSize.{w,v} (forget MonCat.{u}) where + CreatesLimitsOfShape := inferInstance + +@[to_additive] +noncomputable instance forget_createsLimits : + CreatesLimits (forget MonCat.{u}) := CommMonCat.forget_createsLimitsOfSize.{u,u} + end CommMonCat From ed2f6ae5860b941bc054f11d5e402a480667d9e0 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 16 Dec 2024 09:15:29 +0000 Subject: [PATCH 733/829] chore: remove choice from `Multiset.pi` (#19990) Since this doesn't actually complicate the proof or change the statement, this seems harmless. Originated from [this Zulip thread](https://leanprover.zulipchat.com/#narrow/channel/236446-Type-theory/topic/Setless.20proof.20of.20.60Fin2.201.20-.3E.20Fin2.201.20.5Cneq.20Fin2.202.20-.3E.20Fin2.202.60/near/489156076) --- Mathlib/Data/Multiset/Pi.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Data/Multiset/Pi.lean b/Mathlib/Data/Multiset/Pi.lean index f67654f455416..fc8dad8e16e55 100644 --- a/Mathlib/Data/Multiset/Pi.lean +++ b/Mathlib/Data/Multiset/Pi.lean @@ -53,8 +53,8 @@ theorem cons_swap {a a' : α} {b : δ a} {b' : δ a'} {m : Multiset α} {f : ∀ simp only [heq_iff_eq] rintro a'' _ rfl refine hfunext (by rw [Multiset.cons_swap]) fun ha₁ ha₂ _ => ?_ - rcases ne_or_eq a'' a with (h₁ | rfl) - on_goal 1 => rcases eq_or_ne a'' a' with (rfl | h₂) + rcases Decidable.ne_or_eq a'' a with (h₁ | rfl) + on_goal 1 => rcases Decidable.eq_or_ne a'' a' with (rfl | h₂) all_goals simp [*, Pi.cons_same, Pi.cons_ne] @[simp, nolint simpNF] -- Porting note: false positive, this lemma can prove itself From 748cb3923e4ce6ad8d13903e2ed5e286d97b926a Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 16 Dec 2024 09:28:06 +0000 Subject: [PATCH 734/829] chore(SmoothManifoldWithCorners): `Inhabited`->`Nonempty` (#19987) --- Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index 9fcbaff7d4fac..d3b381f0c0511 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -1554,7 +1554,7 @@ lemma LocallyCompactSpace.of_locallyCompact_manifold (I : ModelWithCorners 𝕜 finite-dimensional space. This is the converse to `Manifold.locallyCompact_of_finiteDimensional`. -/ theorem FiniteDimensional.of_locallyCompact_manifold - [CompleteSpace 𝕜] (I : ModelWithCorners 𝕜 E H) [Inhabited M] [LocallyCompactSpace M] : + [CompleteSpace 𝕜] (I : ModelWithCorners 𝕜 E H) [Nonempty M] [LocallyCompactSpace M] : FiniteDimensional 𝕜 E := by have := LocallyCompactSpace.of_locallyCompact_manifold M I exact FiniteDimensional.of_locallyCompactSpace 𝕜 From a1d13da4124ee22aaff652a2e1a9c364002b75ba Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 16 Dec 2024 09:28:07 +0000 Subject: [PATCH 735/829] fix: try with a bash escape to avoid backticks (#19991) [Reported on Zulip](https://leanprover.zulipchat.com/#narrow/channel/428973-nightly-testing/topic/Mergeable.20lean.20testing.20PRs/near/489182290) --- .github/workflows/discover-lean-pr-testing.yml | 9 ++++----- 1 file changed, 4 insertions(+), 5 deletions(-) diff --git a/.github/workflows/discover-lean-pr-testing.yml b/.github/workflows/discover-lean-pr-testing.yml index 71d4a37aef33d..940c2e6d967d8 100644 --- a/.github/workflows/discover-lean-pr-testing.yml +++ b/.github/workflows/discover-lean-pr-testing.yml @@ -129,14 +129,13 @@ jobs: run: | BRANCHES="${{ steps.find-branches.outputs.branches }}" { - printf "msg<> "$GITHUB_OUTPUT" From afdc9d9598aa31f7d273fe58775f577feef0f703 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 16 Dec 2024 09:38:12 +0000 Subject: [PATCH 736/829] chore: rename `DivInvMonoid.Pow` (#19884) This didn't match the naming convention at all. Loogle tells me that `ZPow` is slightly more common than `PowInt`, though both are used. This is an instance, so the name needs no deprecation --- Mathlib/Algebra/Group/Defs.lean | 6 +++--- Mathlib/Data/ZMod/IntUnitsPower.lean | 2 +- MathlibTest/instance_diamonds.lean | 6 +++--- 3 files changed, 7 insertions(+), 7 deletions(-) diff --git a/Mathlib/Algebra/Group/Defs.lean b/Mathlib/Algebra/Group/Defs.lean index 3d069aab663a3..394c147b9bf64 100644 --- a/Mathlib/Algebra/Group/Defs.lean +++ b/Mathlib/Algebra/Group/Defs.lean @@ -858,13 +858,13 @@ class SubNegMonoid (G : Type u) extends AddMonoid G, Neg G, Sub G where attribute [to_additive SubNegMonoid] DivInvMonoid -instance DivInvMonoid.Pow {M} [DivInvMonoid M] : Pow M ℤ := +instance DivInvMonoid.toZPow {M} [DivInvMonoid M] : Pow M ℤ := ⟨fun x n ↦ DivInvMonoid.zpow n x⟩ -instance SubNegMonoid.SMulInt {M} [SubNegMonoid M] : SMul ℤ M := +instance SubNegMonoid.toZSMul {M} [SubNegMonoid M] : SMul ℤ M := ⟨SubNegMonoid.zsmul⟩ -attribute [to_additive existing SubNegMonoid.SMulInt] DivInvMonoid.Pow +attribute [to_additive existing] DivInvMonoid.toZPow /-- A group is called *cyclic* if it is generated by a single element. -/ class IsAddCyclic (G : Type u) [SMul ℤ G] : Prop where diff --git a/Mathlib/Data/ZMod/IntUnitsPower.lean b/Mathlib/Data/ZMod/IntUnitsPower.lean index 4b5cace9f38cd..61cc3341fdfc7 100644 --- a/Mathlib/Data/ZMod/IntUnitsPower.lean +++ b/Mathlib/Data/ZMod/IntUnitsPower.lean @@ -60,7 +60,7 @@ instance Int.instUnitsPow : Pow ℤˣ R where -- The above instances form no typeclass diamonds with the standard power operators -- but we will need `reducible_and_instances` which currently fails https://github.com/leanprover-community/mathlib4/issues/10906 example : Int.instUnitsPow = Monoid.toNatPow := rfl -example : Int.instUnitsPow = DivInvMonoid.Pow := rfl +example : Int.instUnitsPow = DivInvMonoid.toZPow := rfl @[simp] lemma ofMul_uzpow (u : ℤˣ) (r : R) : Additive.ofMul (u ^ r) = r • Additive.ofMul u := rfl diff --git a/MathlibTest/instance_diamonds.lean b/MathlibTest/instance_diamonds.lean index 48cf79849a3b7..02932d2f6192b 100644 --- a/MathlibTest/instance_diamonds.lean +++ b/MathlibTest/instance_diamonds.lean @@ -23,7 +23,7 @@ section SMul open scoped Polynomial -example : (SubNegMonoid.SMulInt : SMul ℤ ℂ) = (Complex.SMul.instSMulRealComplex : SMul ℤ ℂ) := by +example : (SubNegMonoid.toZSMul : SMul ℤ ℂ) = (Complex.SMul.instSMulRealComplex : SMul ℤ ℂ) := by with_reducible_and_instances rfl example : RestrictScalars.module ℝ ℂ ℂ = Complex.instModule := by @@ -38,7 +38,7 @@ example (α β : Type _) [AddMonoid α] [AddMonoid β] : with_reducible_and_instances rfl example (α β : Type _) [SubNegMonoid α] [SubNegMonoid β] : - (Prod.smul : SMul ℤ (α × β)) = SubNegMonoid.SMulInt := by + (Prod.smul : SMul ℤ (α × β)) = SubNegMonoid.toZSMul := by with_reducible_and_instances rfl example (α : Type _) (β : α → Type _) [∀ a, AddMonoid (β a)] : @@ -46,7 +46,7 @@ example (α : Type _) (β : α → Type _) [∀ a, AddMonoid (β a)] : with_reducible_and_instances rfl example (α : Type _) (β : α → Type _) [∀ a, SubNegMonoid (β a)] : - (Pi.instSMul : SMul ℤ (∀ a, β a)) = SubNegMonoid.SMulInt := by + (Pi.instSMul : SMul ℤ (∀ a, β a)) = SubNegMonoid.toZSMul := by with_reducible_and_instances rfl namespace TensorProduct From 896a59b06349b3d9f1ea62d065bded8fd980be9e Mon Sep 17 00:00:00 2001 From: Pim Otte Date: Mon, 16 Dec 2024 10:17:35 +0000 Subject: [PATCH 737/829] feat(Combinatorics/SimpleGraph): the reverse of a cycle is a cycle (#19611) Added the result that a reverse cycle is a cycle. Includes supporting lemmas concerning lists. In preparation for Tutte's theorem. Co-authored-by: Pim Otte --- Mathlib/Combinatorics/SimpleGraph/Path.lean | 9 +++++++++ Mathlib/Combinatorics/SimpleGraph/Walk.lean | 22 +++++++++++++++++++++ Mathlib/Data/List/Basic.lean | 9 +++++++++ Mathlib/Data/List/Nodup.lean | 15 ++++++++++++++ 4 files changed, 55 insertions(+) diff --git a/Mathlib/Combinatorics/SimpleGraph/Path.lean b/Mathlib/Combinatorics/SimpleGraph/Path.lean index ad69b5d9de3a7..34fc605b7283d 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Path.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Path.lean @@ -249,6 +249,15 @@ theorem cons_isCycle_iff {u v : V} (p : G.Walk v u) (h : G.Adj u v) : have : p.support.Nodup → p.edges.Nodup := edges_nodup_of_support_nodup tauto +protected lemma IsCycle.reverse {p : G.Walk u u} (h : p.IsCycle) : p.reverse.IsCycle := by + simp only [Walk.isCycle_def, nodup_tail_support_reverse] at h ⊢ + exact ⟨h.1.reverse, fun h' ↦ h.2.1 (by simp_all [← Walk.length_eq_zero_iff]), h.2.2⟩ + +@[simp] +lemma isCycle_reverse {p : G.Walk u u} : p.reverse.IsCycle ↔ p.IsCycle where + mp h := by simpa using h.reverse + mpr := .reverse + lemma IsPath.tail {p : G.Walk u v} (hp : p.IsPath) (hp' : ¬ p.Nil) : p.tail.IsPath := by rw [Walk.isPath_def] at hp ⊢ rw [← cons_support_tail _ hp', List.nodup_cons] at hp diff --git a/Mathlib/Combinatorics/SimpleGraph/Walk.lean b/Mathlib/Combinatorics/SimpleGraph/Walk.lean index e8975cff53b2c..0b5c0286bb134 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Walk.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Walk.lean @@ -557,6 +557,20 @@ theorem subset_support_append_right {V : Type u} {G : SimpleGraph V} {u v w : V} intro h simp +contextual only [mem_support_append_iff, or_true, imp_true_iff] +lemma getVert_eq_support_get? {u v n} (p : G.Walk u v) (h2 : n ≤ p.length) : + p.getVert n = p.support[n]? := by + match p with + | .nil => simp_all + | .cons h q => + simp only [Walk.support_cons] + by_cases hn : n = 0 + · simp only [hn, getVert_zero, List.length_cons, Nat.zero_lt_succ, List.getElem?_eq_getElem, + List.getElem_cons_zero] + · push_neg at hn + nth_rewrite 2 [show n = (n - 1) + 1 from by omega] + rw [Walk.getVert_cons q h hn, List.getElem?_cons_succ] + exact getVert_eq_support_get? q (Nat.sub_le_of_le_add (Walk.length_cons _ _ ▸ h2)) + theorem coe_support {u v : V} (p : G.Walk u v) : (p.support : Multiset V) = {u} + p.support.tail := by cases p <;> rfl @@ -739,6 +753,14 @@ theorem edges_nodup_of_support_nodup {u v : V} {p : G.Walk u v} (h : p.support.N simp only [edges_cons, support_cons, List.nodup_cons] at h ⊢ exact ⟨fun h' => h.1 (fst_mem_support_of_mem_edges p' h'), ih h.2⟩ +theorem nodup_tail_support_reverse {u : V} {p : G.Walk u u} : + p.reverse.support.tail.Nodup ↔ p.support.tail.Nodup := by + rw [Walk.support_reverse] + refine List.nodup_tail_reverse p.support ?h + rw [← getVert_eq_support_get? _ (by omega), List.getLast?_eq_getElem?, + ← getVert_eq_support_get? _ (by rw [Walk.length_support]; omega)] + aesop + theorem edges_injective {u v : V} : Function.Injective (Walk.edges : G.Walk u v → List (Sym2 V)) | .nil, .nil, _ => rfl | .nil, .cons _ _, h => by simp at h diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 6b8c7c19967e6..9c5ae0bef9266 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -1032,6 +1032,15 @@ theorem cons_get_drop_succ {l : List α} {n} : l.get n :: l.drop (n.1 + 1) = l.drop n.1 := (drop_eq_getElem_cons n.2).symm +lemma drop_length_sub_one {l : List α} (h : l ≠ []) : l.drop (l.length - 1) = [l.getLast h] := by + induction l with + | nil => aesop + | cons a l ih => + by_cases hl : l = [] + · aesop + rw [length_cons, Nat.add_one_sub_one, List.drop_length_cons hl a] + aesop + section TakeI variable [Inhabited α] diff --git a/Mathlib/Data/List/Nodup.lean b/Mathlib/Data/List/Nodup.lean index 1368f8f1890a3..da7ac4e717fb2 100644 --- a/Mathlib/Data/List/Nodup.lean +++ b/Mathlib/Data/List/Nodup.lean @@ -227,6 +227,21 @@ theorem Nodup.filter (p : α → Bool) {l} : Nodup l → Nodup (filter p l) := b theorem nodup_reverse {l : List α} : Nodup (reverse l) ↔ Nodup l := pairwise_reverse.trans <| by simp only [Nodup, Ne, eq_comm] +lemma nodup_tail_reverse (l : List α) (h : l[0]? = l.getLast?) : + Nodup l.reverse.tail ↔ Nodup l.tail := by + induction l with + | nil => simp + | cons a l ih => + by_cases hl : l = [] + · aesop + · simp_all only [List.get?_eq_getElem?, List.tail_reverse, List.nodup_reverse, + List.dropLast_cons_of_ne_nil hl, List.tail_cons] + simp only [length_cons, Nat.zero_lt_succ, getElem?_eq_getElem, getElem_cons_zero, + Nat.add_one_sub_one, Nat.lt_add_one, Option.some.injEq, List.getElem_cons, + show l.length ≠ 0 from by aesop, ↓reduceDIte, getLast?_eq_getElem?] at h + rw [h, show l.Nodup = (l.dropLast ++ [l.getLast hl]).Nodup from by + simp [List.dropLast_eq_take, ← List.drop_length_sub_one], List.nodup_append_comm] + simp [List.getLast_eq_getElem] theorem Nodup.erase_getElem [DecidableEq α] {l : List α} (hl : l.Nodup) (i : Nat) (h : i < l.length) : l.erase l[i] = l.eraseIdx ↑i := by From 337fbe4c381c8f88aafd0c7470f8d4b144b55eef Mon Sep 17 00:00:00 2001 From: Vasily Nesterov Date: Mon, 16 Dec 2024 10:17:36 +0000 Subject: [PATCH 738/829] refactor(Mathlib/Order): remove duplicate lemma `iInf_le'` (#19911) Remove `iInf_le'` which is equal to `iInf_le`. --- Mathlib/Order/CompleteLattice.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/Order/CompleteLattice.lean b/Mathlib/Order/CompleteLattice.lean index ec003a7bc5715..edf6f099fdb2b 100644 --- a/Mathlib/Order/CompleteLattice.lean +++ b/Mathlib/Order/CompleteLattice.lean @@ -623,11 +623,11 @@ theorem le_iSup (f : ι → α) (i : ι) : f i ≤ iSup f := theorem iInf_le (f : ι → α) (i : ι) : iInf f ≤ f i := sInf_le ⟨i, rfl⟩ -theorem le_iSup' (f : ι → α) (i : ι) : f i ≤ iSup f := - le_sSup ⟨i, rfl⟩ +@[deprecated le_iSup (since := "2024-12-13")] +theorem le_iSup' (f : ι → α) (i : ι) : f i ≤ iSup f := le_iSup f i -theorem iInf_le' (f : ι → α) (i : ι) : iInf f ≤ f i := - sInf_le ⟨i, rfl⟩ +@[deprecated iInf_le (since := "2024-12-13")] +theorem iInf_le' (f : ι → α) (i : ι) : iInf f ≤ f i := iInf_le f i theorem isLUB_iSup : IsLUB (range f) (⨆ j, f j) := isLUB_sSup _ From e8b45bfff945c2ac7e80da3fd9d188c2095d3f0f Mon Sep 17 00:00:00 2001 From: Daniel Weber Date: Mon, 16 Dec 2024 11:07:11 +0000 Subject: [PATCH 739/829] feat(Algebra/Polynomial/FieldDivision): factoring polynomials to irreducible monic factors (#19640) Co-authored-by: Daniel Weber <55664973+Command-Master@users.noreply.github.com> --- Mathlib/Algebra/Polynomial/FieldDivision.lean | 14 +++++++ Mathlib/Data/Nat/Squarefree.lean | 6 +-- Mathlib/NumberTheory/KummerDedekind.lean | 2 +- Mathlib/RingTheory/ChainOfDivisors.lean | 2 +- Mathlib/RingTheory/DedekindDomain/Ideal.lean | 2 +- Mathlib/RingTheory/Radical.lean | 2 +- .../UniqueFactorizationDomain/Finite.lean | 4 +- .../Multiplicative.lean | 6 +-- .../Multiplicity.lean | 2 +- .../UniqueFactorizationDomain/Nat.lean | 4 +- .../NormalizedFactors.lean | 37 ++++++++++++------- 11 files changed, 52 insertions(+), 29 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/FieldDivision.lean b/Mathlib/Algebra/Polynomial/FieldDivision.lean index bb20dc5de2f78..3b5ed5bdcc23f 100644 --- a/Mathlib/Algebra/Polynomial/FieldDivision.lean +++ b/Mathlib/Algebra/Polynomial/FieldDivision.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.Polynomial.Derivative import Mathlib.Algebra.Polynomial.Eval.SMul import Mathlib.Algebra.Polynomial.Roots import Mathlib.RingTheory.EuclideanDomain +import Mathlib.RingTheory.UniqueFactorizationDomain.NormalizedFactors /-! # Theory of univariate polynomials @@ -657,6 +658,19 @@ theorem irreducible_iff_lt_natDegree_lt {p : R[X]} (hp0 : p ≠ 0) (hpu : ¬ IsU simp only [IsUnit.dvd_mul_right (isUnit_C.mpr (IsUnit.mk0 (leadingCoeff p)⁻¹ (inv_ne_zero (leadingCoeff_ne_zero.mpr hp0))))] +open UniqueFactorizationMonoid in +/-- +The normalized factors of a polynomial over a field times its leading coefficient give +the polynomial. +-/ +theorem leadingCoeff_mul_prod_normalizedFactors [DecidableEq R] (a : R[X]) : + C a.leadingCoeff * (normalizedFactors a).prod = a := by + by_cases ha : a = 0 + · simp [ha] + rw [prod_normalizedFactors_eq, normalize_apply, coe_normUnit, CommGroupWithZero.coe_normUnit, + mul_comm, mul_assoc, ← map_mul, inv_mul_cancel₀] <;> + simp_all + end Field end Polynomial diff --git a/Mathlib/Data/Nat/Squarefree.lean b/Mathlib/Data/Nat/Squarefree.lean index b1463cae82ce5..c2cdfeb2c6e50 100644 --- a/Mathlib/Data/Nat/Squarefree.lean +++ b/Mathlib/Data/Nat/Squarefree.lean @@ -269,7 +269,7 @@ theorem divisors_filter_squarefree {n : ℕ} (h0 : n ≠ 0) : rw [UniqueFactorizationMonoid.squarefree_iff_nodup_normalizedFactors h0.1] at hsq rw [Multiset.toFinset_subset, Multiset.toFinset_val, hsq.dedup, ← associated_iff_eq, normalizedFactors_mul h0.1 h0.2] - exact ⟨Multiset.subset_of_le (Multiset.le_add_right _ _), normalizedFactors_prod h0.1⟩ + exact ⟨Multiset.subset_of_le (Multiset.le_add_right _ _), prod_normalizedFactors h0.1⟩ · rintro ⟨s, hs, rfl⟩ rw [Finset.mem_powerset, ← Finset.val_le_iff, Multiset.toFinset_val] at hs have hs0 : s.val.prod ≠ 0 := by @@ -278,14 +278,14 @@ theorem divisors_filter_squarefree {n : ℕ} (h0 : n ≠ 0) : apply not_irreducible_zero (irreducible_of_normalized_factor 0 (Multiset.mem_dedup.1 (Multiset.mem_of_le hs con))) - rw [(normalizedFactors_prod h0).symm.dvd_iff_dvd_right] + rw [(prod_normalizedFactors h0).symm.dvd_iff_dvd_right] refine ⟨⟨Multiset.prod_dvd_prod_of_le (le_trans hs (Multiset.dedup_le _)), h0⟩, ?_⟩ have h := UniqueFactorizationMonoid.factors_unique irreducible_of_normalized_factor (fun x hx => irreducible_of_normalized_factor x (Multiset.mem_of_le (le_trans hs (Multiset.dedup_le _)) hx)) - (normalizedFactors_prod hs0) + (prod_normalizedFactors hs0) rw [associated_eq_eq, Multiset.rel_eq] at h rw [UniqueFactorizationMonoid.squarefree_iff_nodup_normalizedFactors hs0, h] apply s.nodup diff --git a/Mathlib/NumberTheory/KummerDedekind.lean b/Mathlib/NumberTheory/KummerDedekind.lean index f3827e812c309..aa1e1913d08e0 100644 --- a/Mathlib/NumberTheory/KummerDedekind.lean +++ b/Mathlib/NumberTheory/KummerDedekind.lean @@ -324,7 +324,7 @@ theorem Ideal.irreducible_map_of_irreducible_minpoly (hI : IsMaximal I) (hI' : I simp [normalizedFactors_irreducible hf] suffices ∃ y, normalizedFactors (I.map (algebraMap R S)) = {y} by obtain ⟨y, hy⟩ := this - have h := normalizedFactors_prod (show I.map (algebraMap R S) ≠ 0 by + have h := prod_normalizedFactors (show I.map (algebraMap R S) ≠ 0 by rwa [← bot_eq_zero, Ne, map_eq_bot_iff_of_injective (NoZeroSMulDivisors.algebraMap_injective R S)]) rw [associated_iff_eq, hy, Multiset.prod_singleton] at h diff --git a/Mathlib/RingTheory/ChainOfDivisors.lean b/Mathlib/RingTheory/ChainOfDivisors.lean index 68af812321e63..40a9ea9b28e38 100644 --- a/Mathlib/RingTheory/ChainOfDivisors.lean +++ b/Mathlib/RingTheory/ChainOfDivisors.lean @@ -151,7 +151,7 @@ theorem element_of_chain_eq_pow_second_of_chain {q r : Associates M} {n : ℕ} ( eq_second_of_chain_of_prime_dvd hn h₁ (@fun r' => h₂) (prime_of_normalized_factor b hb) hr (dvd_of_mem_normalizedFactors hb) have H : r = c 1 ^ i := by - have := UniqueFactorizationMonoid.normalizedFactors_prod (ne_zero_of_dvd_ne_zero hq hr) + have := UniqueFactorizationMonoid.prod_normalizedFactors (ne_zero_of_dvd_ne_zero hq hr) rw [associated_iff_eq, hi, Multiset.prod_replicate] at this rw [this] refine ⟨⟨i, ?_⟩, H⟩ diff --git a/Mathlib/RingTheory/DedekindDomain/Ideal.lean b/Mathlib/RingTheory/DedekindDomain/Ideal.lean index acc74b557241b..403dfb54a459e 100644 --- a/Mathlib/RingTheory/DedekindDomain/Ideal.lean +++ b/Mathlib/RingTheory/DedekindDomain/Ideal.lean @@ -856,7 +856,7 @@ variable {T : Type*} [CommRing T] [IsDedekindDomain T] {I J : Ideal T} open Multiset UniqueFactorizationMonoid Ideal theorem prod_normalizedFactors_eq_self (hI : I ≠ ⊥) : (normalizedFactors I).prod = I := - associated_iff_eq.1 (normalizedFactors_prod hI) + associated_iff_eq.1 (prod_normalizedFactors hI) theorem count_le_of_ideal_ge [DecidableEq (Ideal T)] {I J : Ideal T} (h : I ≤ J) (hI : I ≠ ⊥) (K : Ideal T) : diff --git a/Mathlib/RingTheory/Radical.lean b/Mathlib/RingTheory/Radical.lean index cea98e424f69b..1516e4ffea8f0 100644 --- a/Mathlib/RingTheory/Radical.lean +++ b/Mathlib/RingTheory/Radical.lean @@ -105,7 +105,7 @@ theorem radical_dvd_self (a : M) : radical a ∣ a := by by_cases ha : a = 0 · rw [ha] apply dvd_zero - · rw [radical, ← Finset.prod_val, ← (normalizedFactors_prod ha).dvd_iff_dvd_right] + · rw [radical, ← Finset.prod_val, ← (prod_normalizedFactors ha).dvd_iff_dvd_right] apply Multiset.prod_dvd_prod_of_le rw [primeFactors, Multiset.toFinset_val] apply Multiset.dedup_le diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Finite.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Finite.lean index bd5cd5c1e9819..a2a89c9d76ef2 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain/Finite.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Finite.lean @@ -40,13 +40,13 @@ noncomputable def fintypeSubtypeDvd {M : Type*} [CancelCommMonoidWithZero M] · rintro ⟨s, hs, rfl⟩ show (s.snd : M) * s.fst.prod ∣ y rw [(unit_associated_one.mul_right s.fst.prod).dvd_iff_dvd_left, one_mul, - ← (normalizedFactors_prod hy).dvd_iff_dvd_right] + ← (prod_normalizedFactors hy).dvd_iff_dvd_right] exact Multiset.prod_dvd_prod_of_le hs · rintro (h : x ∣ y) have hx : x ≠ 0 := by refine mt (fun hx => ?_) hy rwa [hx, zero_dvd_iff] at h - obtain ⟨u, hu⟩ := normalizedFactors_prod hx + obtain ⟨u, hu⟩ := prod_normalizedFactors hx refine ⟨⟨normalizedFactors x, u⟩, ?_, (mul_comm _ _).trans hu⟩ exact (dvd_iff_normalizedFactors_le_normalizedFactors hx hy).mp h diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicative.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicative.lean index a8a04d3a9ad0f..8c796646ffad7 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicative.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicative.lean @@ -78,7 +78,7 @@ theorem induction_on_coprime {P : α → Prop} (a : α) (h0 : P 0) (h1 : ∀ {x} · rwa [ha0] haveI : Nontrivial α := ⟨⟨_, _, ha0⟩⟩ letI : NormalizationMonoid α := UniqueFactorizationMonoid.normalizationMonoid - refine P_of_associated (normalizedFactors_prod ha0) ?_ + refine P_of_associated (prod_normalizedFactors ha0) ?_ rw [← (normalizedFactors a).map_id, Finset.prod_multiset_map_count] refine induction_on_prime_power _ _ ?_ ?_ @h1 @hpr @hcp <;> simp only [Multiset.mem_toFinset] · apply prime_of_normalized_factor @@ -131,8 +131,8 @@ theorem multiplicative_of_coprime (f : α → β) (a b : α) (h0 : f 0 = 0) p ^ (normalizedFactors a).count p) * f (∏ p ∈ (normalizedFactors a).toFinset ∪ (normalizedFactors b).toFinset, p ^ (normalizedFactors b).count p) by - obtain ⟨ua, a_eq⟩ := normalizedFactors_prod ha0 - obtain ⟨ub, b_eq⟩ := normalizedFactors_prod hb0 + obtain ⟨ua, a_eq⟩ := prod_normalizedFactors ha0 + obtain ⟨ub, b_eq⟩ := prod_normalizedFactors hb0 rw [← a_eq, ← b_eq, mul_right_comm (Multiset.prod (normalizedFactors a)) ua (Multiset.prod (normalizedFactors b) * ub), h1 ua.isUnit, h1 ub.isUnit, h1 ua.isUnit, ← mul_assoc, h1 ub.isUnit, mul_right_comm _ (f ua), ← mul_assoc] diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean index a6504ae1d7bda..9768ee91a10df 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean @@ -75,7 +75,7 @@ theorem le_emultiplicity_iff_replicate_le_normalizedFactors {a b : R} {n : ℕ} apply Dvd.intro _ rfl · rw [Multiset.le_iff_exists_add] rintro ⟨u, hu⟩ - rw [← (normalizedFactors_prod hb).dvd_iff_dvd_right, hu, prod_add, prod_replicate] + rw [← (prod_normalizedFactors hb).dvd_iff_dvd_right, hu, prod_add, prod_replicate] exact (Associated.pow_pow <| associated_normalize a).dvd.trans (Dvd.intro u.prod rfl) /-- The multiplicity of an irreducible factor of a nonzero element is exactly the number of times diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean index d8e8cd87f6ef9..0955b414f81e9 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean @@ -44,7 +44,7 @@ lemma factors_eq : ∀ n : ℕ, normalizedFactors n = n.primeFactorsList rw [← Multiset.rel_eq, ← associated_eq_eq] apply UniqueFactorizationMonoid.factors_unique irreducible_of_normalized_factor _ · rw [Multiset.prod_coe, Nat.prod_primeFactorsList n.succ_ne_zero] - exact normalizedFactors_prod n.succ_ne_zero + exact prod_normalizedFactors n.succ_ne_zero · intro x hx rw [Nat.irreducible_iff_prime, ← Nat.prime_iff] exact Nat.prime_of_mem_primeFactorsList hx @@ -53,7 +53,7 @@ lemma factors_multiset_prod_of_irreducible {s : Multiset ℕ} (h : ∀ x : ℕ, normalizedFactors s.prod = s := by rw [← Multiset.rel_eq, ← associated_eq_eq] apply UniqueFactorizationMonoid.factors_unique irreducible_of_normalized_factor h - (normalizedFactors_prod _) + (prod_normalizedFactors _) rw [Ne, Multiset.prod_eq_zero_iff] exact fun con ↦ not_irreducible_zero (h 0 con) diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/NormalizedFactors.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/NormalizedFactors.lean index 4db1bb549e738..5f137806b6aa5 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain/NormalizedFactors.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/NormalizedFactors.lean @@ -39,7 +39,7 @@ theorem factors_eq_normalizedFactors {M : Type*} [CancelCommMonoidWithZero M] ext p exact normalize_eq p -theorem normalizedFactors_prod {a : α} (ane0 : a ≠ 0) : +theorem prod_normalizedFactors {a : α} (ane0 : a ≠ 0) : Associated (normalizedFactors a).prod a := by rw [normalizedFactors, factors, dif_neg ane0] refine Associated.trans ?_ (Classical.choose_spec (exists_prime_factors a ane0)).2 @@ -49,6 +49,15 @@ theorem normalizedFactors_prod {a : α} (ane0 : a ≠ 0) : ext rw [Function.comp_apply, Associates.mk_normalize] +@[deprecated (since := "2024-12-04")] +alias normalizedFactors_prod := prod_normalizedFactors + +theorem prod_normalizedFactors_eq {a : α} (ane0 : a ≠ 0) : + (normalizedFactors a).prod = normalize a := by + trans normalize (normalizedFactors a).prod + · rw [normalizedFactors, ← map_multiset_prod, normalize_idem] + · exact normalize_eq_normalize_iff.mpr (dvd_dvd_iff_associated.mpr (prod_normalizedFactors ane0)) + theorem prime_of_normalized_factor {a : α} : ∀ x : α, x ∈ normalizedFactors a → Prime x := by rw [normalizedFactors, factors] split_ifs with ane0; · simp @@ -71,7 +80,7 @@ theorem normalize_normalized_factor {a : α} : theorem normalizedFactors_irreducible {a : α} (ha : Irreducible a) : normalizedFactors a = {normalize a} := by obtain ⟨p, a_assoc, hp⟩ := - prime_factors_irreducible ha ⟨prime_of_normalized_factor, normalizedFactors_prod ha.ne_zero⟩ + prime_factors_irreducible ha ⟨prime_of_normalized_factor, prod_normalizedFactors ha.ne_zero⟩ have p_mem : p ∈ normalizedFactors a := by rw [hp] exact Multiset.mem_singleton_self _ @@ -96,11 +105,11 @@ theorem exists_mem_normalizedFactors_of_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irr irreducible_of_normalized_factor (Associated.symm <| calc - Multiset.prod (normalizedFactors a) ~ᵤ a := normalizedFactors_prod ha0 + Multiset.prod (normalizedFactors a) ~ᵤ a := prod_normalizedFactors ha0 _ = p * b := hb _ ~ᵤ Multiset.prod (p ::ₘ normalizedFactors b) := by rw [Multiset.prod_cons] - exact (normalizedFactors_prod hb0).symm.mul_left _ + exact (prod_normalizedFactors hb0).symm.mul_left _ ) Multiset.exists_mem_of_rel_of_mem this (by simp) @@ -124,7 +133,7 @@ theorem normalizedFactors_one : normalizedFactors (1 : α) = 0 := by · intro x hx exfalso apply Multiset.not_mem_zero x hx - · apply normalizedFactors_prod one_ne_zero + · apply prod_normalizedFactors one_ne_zero @[simp] theorem normalizedFactors_mul {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : @@ -145,8 +154,8 @@ theorem normalizedFactors_mul {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : · exact irreducible_of_normalized_factor · rw [Multiset.prod_add] exact - ((normalizedFactors_prod hx).mul_mul (normalizedFactors_prod hy)).trans - (normalizedFactors_prod (mul_ne_zero hx hy)).symm + ((prod_normalizedFactors hx).mul_mul (prod_normalizedFactors hy)).trans + (prod_normalizedFactors (mul_ne_zero hx hy)).symm @[simp] theorem normalizedFactors_pow {x : α} (n : ℕ) : @@ -181,8 +190,8 @@ theorem dvd_iff_normalizedFactors_le_normalizedFactors {x y : α} (hx : x ≠ 0) constructor · rintro ⟨c, rfl⟩ simp [hx, right_ne_zero_of_mul hy] - · rw [← (normalizedFactors_prod hx).dvd_iff_dvd_left, ← - (normalizedFactors_prod hy).dvd_iff_dvd_right] + · rw [← (prod_normalizedFactors hx).dvd_iff_dvd_left, ← + (prod_normalizedFactors hy).dvd_iff_dvd_right] apply Multiset.prod_dvd_prod_of_le theorem _root_.Associated.normalizedFactors_eq {a b : α} (h : Associated a b) : @@ -195,7 +204,7 @@ theorem _root_.Associated.normalizedFactors_eq {a b : α} (h : Associated a b) : theorem associated_iff_normalizedFactors_eq_normalizedFactors {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : x ~ᵤ y ↔ normalizedFactors x = normalizedFactors y := ⟨Associated.normalizedFactors_eq, fun h => - (normalizedFactors_prod hx).symm.trans (_root_.trans (by rw [h]) (normalizedFactors_prod hy))⟩ + (prod_normalizedFactors hx).symm.trans (_root_.trans (by rw [h]) (prod_normalizedFactors hy))⟩ theorem normalizedFactors_of_irreducible_pow {p : α} (hp : Irreducible p) (k : ℕ) : normalizedFactors (p ^ k) = Multiset.replicate k (normalize p) := by @@ -208,7 +217,7 @@ theorem dvd_of_mem_normalizedFactors {a p : α} (H : p ∈ normalizedFactors a) by_cases hcases : a = 0 · rw [hcases] exact dvd_zero p - · exact dvd_trans (Multiset.dvd_prod H) (Associated.dvd (normalizedFactors_prod hcases)) + · exact dvd_trans (Multiset.dvd_prod H) (Associated.dvd (prod_normalizedFactors hcases)) theorem mem_normalizedFactors_iff [Subsingleton αˣ] {p x : α} (hx : x ≠ 0) : p ∈ normalizedFactors x ↔ Prime p ∧ p ∣ x := by @@ -223,7 +232,7 @@ theorem mem_normalizedFactors_iff [Subsingleton αˣ] {p x : α} (hx : x ≠ 0) theorem exists_associated_prime_pow_of_unique_normalized_factor {p r : α} (h : ∀ {m}, m ∈ normalizedFactors r → m = p) (hr : r ≠ 0) : ∃ i : ℕ, Associated (p ^ i) r := by use (normalizedFactors r).card - have := UniqueFactorizationMonoid.normalizedFactors_prod hr + have := UniqueFactorizationMonoid.prod_normalizedFactors hr rwa [Multiset.eq_replicate_of_mem fun b => h, Multiset.prod_replicate] at this theorem normalizedFactors_prod_of_prime [Subsingleton αˣ] {m : Multiset α} @@ -235,7 +244,7 @@ theorem normalizedFactors_prod_of_prime [Subsingleton αˣ] {m : Multiset α} simp · simpa only [← Multiset.rel_eq, ← associated_eq_eq] using prime_factors_unique prime_of_normalized_factor h - (normalizedFactors_prod (m.prod_ne_zero_of_prime h)) + (prod_normalizedFactors (m.prod_ne_zero_of_prime h)) theorem mem_normalizedFactors_eq_of_associated {a b c : α} (ha : a ∈ normalizedFactors c) (hb : b ∈ normalizedFactors c) (h : Associated a b) : a = b := by @@ -324,6 +333,6 @@ protected noncomputable def normalizationMonoid : NormalizationMonoid α := rfl rw [if_neg hx, ← mkMonoidHom_apply, MonoidHom.map_multiset_prod, map_map, h, map_id, ← associated_iff_eq] - apply normalizedFactors_prod hx) + apply prod_normalizedFactors hx) end UniqueFactorizationMonoid From 95e4d18407c257812c48cf37309b701918c5bf1d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 16 Dec 2024 11:23:10 +0000 Subject: [PATCH 740/829] chore(SetTheory/Game/PGame): deprecate primed lemmas (#18469) We perform various renames with the intent of doing away with some badly named or dubiously useful lemmas. --- Mathlib/SetTheory/Game/Basic.lean | 5 +--- Mathlib/SetTheory/Game/Impartial.lean | 4 +-- Mathlib/SetTheory/Game/Nim.lean | 22 ++++++++++----- Mathlib/SetTheory/Game/PGame.lean | 27 ++++++++++++------- Mathlib/SetTheory/Surreal/Multiplication.lean | 11 ++++---- 5 files changed, 42 insertions(+), 27 deletions(-) diff --git a/Mathlib/SetTheory/Game/Basic.lean b/Mathlib/SetTheory/Game/Basic.lean index 155edf8c3a215..be2e92919662c 100644 --- a/Mathlib/SetTheory/Game/Basic.lean +++ b/Mathlib/SetTheory/Game/Basic.lean @@ -790,10 +790,7 @@ def mulOption (x y : PGame) (i : LeftMoves x) (j : LeftMoves y) : PGame := the first kind. -/ lemma mulOption_neg_neg {x} (y) {i j} : mulOption x y i j = mulOption x (-(-y)) i (toLeftMovesNeg <| toRightMovesNeg j) := by - dsimp only [mulOption] - congr 2 - · rw [neg_neg] - iterate 2 rw [moveLeft_neg, moveRight_neg, neg_neg] + simp [mulOption] /-- The left options of `x * y` agree with that of `y * x` up to equivalence. -/ lemma mulOption_symm (x y) {i j} : ⟦mulOption x y i j⟧ = (⟦mulOption y x j i⟧ : Game) := by diff --git a/Mathlib/SetTheory/Game/Impartial.lean b/Mathlib/SetTheory/Game/Impartial.lean index fea623654ef44..e06e5beb85b35 100644 --- a/Mathlib/SetTheory/Game/Impartial.lean +++ b/Mathlib/SetTheory/Game/Impartial.lean @@ -94,9 +94,9 @@ instance impartial_neg (G : PGame) [G.Impartial] : (-G).Impartial := by refine ⟨?_, fun i => ?_, fun i => ?_⟩ · rw [neg_neg] exact Equiv.symm (neg_equiv_self G) - · rw [moveLeft_neg'] + · rw [moveLeft_neg] exact impartial_neg _ - · rw [moveRight_neg'] + · rw [moveRight_neg] exact impartial_neg _ termination_by G diff --git a/Mathlib/SetTheory/Game/Nim.lean b/Mathlib/SetTheory/Game/Nim.lean index 57ac08225b67f..dba76e6378507 100644 --- a/Mathlib/SetTheory/Game/Nim.lean +++ b/Mathlib/SetTheory/Game/Nim.lean @@ -84,16 +84,26 @@ theorem toRightMovesNim_symm_lt {o : Ordinal} (i : (nim o).RightMoves) : (toRightMovesNim.symm i).prop @[simp] -theorem moveLeft_nim' {o : Ordinal} (i) : (nim o).moveLeft i = nim (toLeftMovesNim.symm i).val := +theorem moveLeft_nim {o : Ordinal} (i) : (nim o).moveLeft i = nim (toLeftMovesNim.symm i).val := (congr_heq (moveLeft_nim_hEq o).symm (cast_heq _ i)).symm -theorem moveLeft_nim {o : Ordinal} (i) : (nim o).moveLeft (toLeftMovesNim i) = nim i := by simp +@[deprecated moveLeft_nim (since := "2024-10-30")] +alias moveLeft_nim' := moveLeft_nim + +theorem moveLeft_toLeftMovesNim {o : Ordinal} (i) : + (nim o).moveLeft (toLeftMovesNim i) = nim i := by + simp @[simp] -theorem moveRight_nim' {o : Ordinal} (i) : (nim o).moveRight i = nim (toRightMovesNim.symm i).val := +theorem moveRight_nim {o : Ordinal} (i) : (nim o).moveRight i = nim (toRightMovesNim.symm i).val := (congr_heq (moveRight_nim_hEq o).symm (cast_heq _ i)).symm -theorem moveRight_nim {o : Ordinal} (i) : (nim o).moveRight (toRightMovesNim i) = nim i := by simp +@[deprecated moveRight_nim (since := "2024-10-30")] +alias moveRight_nim' := moveRight_nim + +theorem moveRight_toRightMovesNim {o : Ordinal} (i) : + (nim o).moveRight (toRightMovesNim i) = nim i := by + simp /-- A recursion principle for left moves of a nim game. -/ @[elab_as_elim] @@ -268,7 +278,7 @@ theorem equiv_nim_grundyValue (G : PGame.{u}) [G.Impartial] : · rw [add_moveLeft_inr, ← Impartial.exists_left_move_equiv_iff_fuzzy_zero] obtain ⟨j, hj⟩ := exists_grundyValue_moveLeft_of_lt <| toLeftMovesNim_symm_lt i use toLeftMovesAdd (Sum.inl j) - rw [add_moveLeft_inl, moveLeft_nim'] + rw [add_moveLeft_inl, moveLeft_nim] exact Equiv.trans (add_congr_left (equiv_nim_grundyValue _)) (hj ▸ Impartial.add_self _) termination_by G @@ -329,7 +339,7 @@ theorem exists_grundyValue_moveRight_of_lt {G : PGame} [G.Impartial] {o : Nimber rw [← grundyValue_neg] at h obtain ⟨i, hi⟩ := exists_grundyValue_moveLeft_of_lt h use toLeftMovesNeg.symm i - rwa [← grundyValue_neg, ← moveLeft_neg'] + rwa [← grundyValue_neg, ← moveLeft_neg] theorem grundyValue_le_of_forall_moveRight {G : PGame} [G.Impartial] {o : Nimber} (h : ∀ i, grundyValue (G.moveRight i) ≠ o) : G.grundyValue ≤ o := by diff --git a/Mathlib/SetTheory/Game/PGame.lean b/Mathlib/SetTheory/Game/PGame.lean index 91a9c5e7af475..10fefac3ec058 100644 --- a/Mathlib/SetTheory/Game/PGame.lean +++ b/Mathlib/SetTheory/Game/PGame.lean @@ -1310,34 +1310,43 @@ between them. -/ def toRightMovesNeg {x : PGame} : x.LeftMoves ≃ (-x).RightMoves := Equiv.cast (rightMoves_neg x).symm -theorem moveLeft_neg {x : PGame} (i) : (-x).moveLeft (toLeftMovesNeg i) = -x.moveRight i := by - cases x - rfl - @[simp] -theorem moveLeft_neg' {x : PGame} (i) : (-x).moveLeft i = -x.moveRight (toLeftMovesNeg.symm i) := by +theorem moveLeft_neg {x : PGame} (i) : + (-x).moveLeft i = -x.moveRight (toLeftMovesNeg.symm i) := by cases x rfl -theorem moveRight_neg {x : PGame} (i) : (-x).moveRight (toRightMovesNeg i) = -x.moveLeft i := by - cases x - rfl +@[deprecated moveLeft_neg (since := "2024-10-30")] +alias moveLeft_neg' := moveLeft_neg + +theorem moveLeft_neg_toLeftMovesNeg {x : PGame} (i) : + (-x).moveLeft (toLeftMovesNeg i) = -x.moveRight i := by simp @[simp] -theorem moveRight_neg' {x : PGame} (i) : +theorem moveRight_neg {x : PGame} (i) : (-x).moveRight i = -x.moveLeft (toRightMovesNeg.symm i) := by cases x rfl +@[deprecated moveRight_neg (since := "2024-10-30")] +alias moveRight_neg' := moveRight_neg + +theorem moveRight_neg_toRightMovesNeg {x : PGame} (i) : + (-x).moveRight (toRightMovesNeg i) = -x.moveLeft i := by simp + +@[deprecated moveRight_neg (since := "2024-10-30")] theorem moveLeft_neg_symm {x : PGame} (i) : x.moveLeft (toRightMovesNeg.symm i) = -(-x).moveRight i := by simp +@[deprecated moveRight_neg (since := "2024-10-30")] theorem moveLeft_neg_symm' {x : PGame} (i) : x.moveLeft i = -(-x).moveRight (toRightMovesNeg i) := by simp +@[deprecated moveLeft_neg (since := "2024-10-30")] theorem moveRight_neg_symm {x : PGame} (i) : x.moveRight (toLeftMovesNeg.symm i) = -(-x).moveLeft i := by simp +@[deprecated moveLeft_neg (since := "2024-10-30")] theorem moveRight_neg_symm' {x : PGame} (i) : x.moveRight i = -(-x).moveLeft (toLeftMovesNeg i) := by simp diff --git a/Mathlib/SetTheory/Surreal/Multiplication.lean b/Mathlib/SetTheory/Surreal/Multiplication.lean index 581ee2b7acc94..1a89e79034781 100644 --- a/Mathlib/SetTheory/Surreal/Multiplication.lean +++ b/Mathlib/SetTheory/Surreal/Multiplication.lean @@ -119,7 +119,7 @@ lemma P2_neg_right : P2 x₁ x₂ y ↔ P2 x₁ x₂ (-y) := by rw [P2, P2, quot_mul_neg, quot_mul_neg, neg_inj] lemma P4_neg_left : P4 x₁ x₂ y ↔ P4 (-x₂) (-x₁) y := by - simp_rw [P4, PGame.neg_lt_neg_iff, moveLeft_neg', ← P3_neg] + simp_rw [P4, PGame.neg_lt_neg_iff, moveLeft_neg, ← P3_neg] lemma P4_neg_right : P4 x₁ x₂ y ↔ P4 x₁ x₂ (-y) := by rw [P4, P4, neg_neg, and_comm] @@ -230,7 +230,7 @@ lemma ih1_swap (ih : ∀ a, ArgsRel a (Args.P1 x y) → P124 a) : IH1 y x := ih1 lemma P3_of_ih (hy : Numeric y) (ihyx : IH1 y x) (i k l) : P3 (x.moveLeft i) x (y.moveLeft k) (-(-y).moveLeft l) := P3_comm.2 <| ((ihyx (IsOption.moveLeft k) (isOption_neg.1 <| .moveLeft l) <| Or.inl rfl).2 - (by rw [← moveRight_neg_symm]; apply hy.left_lt_right)).1 i + (by rw [moveLeft_neg, neg_neg]; apply hy.left_lt_right)).1 i lemma P24_of_ih (ihxy : IH1 x y) (i j) : P24 (x.moveLeft i) (x.moveLeft j) y := ihxy (IsOption.moveLeft i) (IsOption.moveLeft j) (Or.inl rfl) @@ -413,9 +413,8 @@ theorem P3_of_lt {y₁ y₂} (h : ∀ i, IH3 x₁ (x₂.moveLeft i) x₂ y₁ y P3 x₁ x₂ y₁ y₂ := by obtain (⟨i,hi⟩|⟨i,hi⟩) := lf_iff_exists_le.1 (lf_of_lt hl) · exact P3_of_le_left i (h i) hi - · exact P3_neg.2 <| P3_of_le_left _ (hs _) <| by - rw [moveLeft_neg] - exact neg_le_neg (le_iff_game_le.1 hi) + · apply P3_neg.2 <| P3_of_le_left _ (hs (toLeftMovesNeg i)) _ + simpa /-- The main chunk of Theorem 8 in [Conway2001] / Theorem 3.8 in [SchleicherStoll]. -/ theorem main (a : Args) : a.Numeric → P124 a := by @@ -488,7 +487,7 @@ theorem P3_of_lt_of_lt (hx₁ : x₁.Numeric) (hx₂ : x₂.Numeric) (hy₁ : y · have hi := hx₁.neg.moveLeft i exact ⟨(P24 hx₂.neg hi hy₁).1, (P24 hx₂.neg hi hy₂).1, P3_comm.2 <| ((P24 hy₁ hy₂ hx₁).2 hy).2 _, by - rw [moveLeft_neg', ← P3_neg, neg_lt_neg_iff] + rw [moveLeft_neg, ← P3_neg, neg_lt_neg_iff] exact ih _ (fst <| IsOption.moveRight _) (hx₁.moveRight _) hx₂⟩ theorem Numeric.mul_pos (hx₁ : x₁.Numeric) (hx₂ : x₂.Numeric) (hp₁ : 0 < x₁) (hp₂ : 0 < x₂) : From 30c7d7b4cc3d048d99c66bac501aaeaa43a1e70a Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 16 Dec 2024 13:05:35 +0000 Subject: [PATCH 741/829] feat(AlgebraicGeometry/EllipticCurve/Projective): implement group law for projective coordinates (#8485) Completes the proof of the group law in projective coordinates. Co-authored-by: Kevin Buzzard --- .../EllipticCurve/Group.lean | 72 +++++++++++++------ .../EllipticCurve/Projective.lean | 3 +- 2 files changed, 52 insertions(+), 23 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean index faa80b6bd843f..a54aea86aeb89 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: David Kurniadi Angdinata -/ import Mathlib.AlgebraicGeometry.EllipticCurve.Jacobian +import Mathlib.AlgebraicGeometry.EllipticCurve.Projective import Mathlib.LinearAlgebra.FreeModule.Norm import Mathlib.RingTheory.ClassGroup import Mathlib.RingTheory.Polynomial.UniqueFactorization @@ -11,31 +12,35 @@ import Mathlib.RingTheory.Polynomial.UniqueFactorization /-! # Group law on Weierstrass curves -This file proves that the nonsingular rational points on a Weierstrass curve forms an abelian group -under the geometric group law defined in `Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean` and -in `Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean`. +This file proves that the nonsingular rational points on a Weierstrass curve form an abelian group +under the geometric group law defined in `Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean`, in +`Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean`, and in +`Mathlib/AlgebraicGeometry/EllipticCurve/Projective.lean`. ## Mathematical background -Let `W` be a Weierstrass curve over a field `F` given by a Weierstrass equation $W(X, Y) = 0$ in -affine coordinates. As in `Mathlib.AlgebraicGeometry.EllipticCurve.Affine`, the set of nonsingular -rational points `W⟮F⟯` of `W` consist of the unique point at infinity $0$ and nonsingular affine -points $(x, y)$. With this description, there is an addition-preserving injection between `W⟮F⟯` -and the ideal class group of the coordinate ring $F[W] := F[X, Y] / \langle W(X, Y)\rangle$ of `W`. -This is defined by mapping the point at infinity $0$ to the trivial ideal class and an affine point -$(x, y)$ to the ideal class of the invertible fractional ideal $\langle X - x, Y - y\rangle$. -Proving that this is well-defined and preserves addition reduce to checking several equalities of -integral ideals, which is done in `WeierstrassCurve.Affine.CoordinateRing.XYIdeal_neg_mul` and in -`WeierstrassCurve.Affine.CoordinateRing.XYIdeal_mul_XYIdeal` via explicit ideal computations. -Now $F[W]$ is a free rank two $F[X]$-algebra with basis $\{1, Y\}$, so every element of $F[W]$ is -of the form $p + qY$ for some $p, q \in F[X]$, and there is an algebra norm $N : F[W] \to F[X]$. -Injectivity can then be shown by computing the degree of such a norm $N(p + qY)$ in two different -ways, which is done in `WeierstrassCurve.Affine.CoordinateRing.degree_norm_smul_basis` and in the -auxiliary lemmas in the proof of `WeierstrassCurve.Affine.Point.instAddCommGroup`. +Let `W` be a Weierstrass curve over a field `F` given by a Weierstrass equation `W(X, Y) = 0` in +affine coordinates. As in `Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean`, the set of +nonsingular rational points `W⟮F⟯` of `W` consist of the unique point at infinity `𝓞` and +nonsingular affine points `(x, y)`. With this description, there is an addition-preserving injection +between `W⟮F⟯` and the ideal class group of the coordinate ring `F[W] := F[X, Y] / ⟨W(X, Y)⟩` of +`W`. This is defined by mapping the point at infinity `𝓞` to the trivial ideal class and an affine +point `(x, y)` to the ideal class of the invertible fractional ideal `⟨X - x, Y - y⟩`. Proving that +this is well-defined and preserves addition reduce to checking several equalities of integral +ideals, which is done in `WeierstrassCurve.Affine.CoordinateRing.XYIdeal_neg_mul` and in +`WeierstrassCurve.Affine.CoordinateRing.XYIdeal_mul_XYIdeal` via explicit ideal computations. Now +`F[W]` is a free rank two `F[X]`-algebra with basis `{1, Y}`, so every element of `F[W]` is of the +form `p + qY` for some `p, q ∈ F[X]`, and there is an algebra norm `N : F[W] → F[X]`. Injectivity +can then be shown by computing the degree of such a norm `N(p + qY)` in two different ways, which is +done in `WeierstrassCurve.Affine.CoordinateRing.degree_norm_smul_basis` and in the auxiliary lemmas +in the proof of `WeierstrassCurve.Affine.Point.instAddCommGroup`. When `W` is given in Jacobian coordinates, `WeierstrassCurve.Jacobian.Point.toAffineAddEquiv` pulls back the group law on `WeierstrassCurve.Affine.Point` to `WeierstrassCurve.Jacobian.Point`. +When `W` is given in projective coordinates, `WeierstrassCurve.Projective.Point.toAffineAddEquiv` +pulls back the group law on `WeierstrassCurve.Affine.Point` to `WeierstrassCurve.Projective.Point`. + ## Main definitions * `WeierstrassCurve.Affine.CoordinateRing`: the coordinate ring `F[W]` of a Weierstrass curve `W`. @@ -49,8 +54,10 @@ back the group law on `WeierstrassCurve.Affine.Point` to `WeierstrassCurve.Jacob element in the coordinate ring of an affine Weierstrass curve in terms of the power basis. * `WeierstrassCurve.Affine.Point.instAddCommGroup`: the type of nonsingular rational points on an affine Weierstrass curve forms an abelian group under addition. - * `WeierstrassCurve.Jacobian.Point.instAddCommGroup`: the type of nonsingular rational - points on a Jacobian Weierstrass curve forms an abelian group under addition. + * `WeierstrassCurve.Jacobian.Point.instAddCommGroup`: the type of nonsingular rational points on a + Jacobian Weierstrass curve forms an abelian group under addition. + * `WeierstrassCurve.Projective.Point.instAddCommGroup`: the type of nonsingular rational points on + a projective Weierstrass curve forms an abelian group under addition. ## References @@ -61,8 +68,9 @@ https://drops.dagstuhl.de/storage/00lipics/lipics-vol268-itp2023/LIPIcs.ITP.2023 elliptic curve, group law, class group -/ -open Ideal nonZeroDivisors Polynomial -open scoped Polynomial.Bivariate +open Ideal Polynomial + +open scoped nonZeroDivisors Polynomial.Bivariate local macro "C_simp" : tactic => `(tactic| simp only [map_ofNat, C_0, C_1, C_neg, C_add, C_sub, C_mul, C_pow]) @@ -562,6 +570,26 @@ end Point end WeierstrassCurve.Affine +namespace WeierstrassCurve.Projective.Point + +/-! ## Weierstrass curves in projective coordinates -/ + +variable {F : Type u} [Field F] {W : Projective F} + +noncomputable instance : AddCommGroup W.Point where + nsmul := nsmulRec + zsmul := zsmulRec + zero_add _ := (toAffineAddEquiv W).injective <| by + simp only [map_add, toAffineAddEquiv_apply, toAffineLift_zero, zero_add] + add_zero _ := (toAffineAddEquiv W).injective <| by + simp only [map_add, toAffineAddEquiv_apply, toAffineLift_zero, add_zero] + neg_add_cancel P := (toAffineAddEquiv W).injective <| by + simp only [map_add, toAffineAddEquiv_apply, toAffineLift_neg, neg_add_cancel, toAffineLift_zero] + add_comm _ _ := (toAffineAddEquiv W).injective <| by simp only [map_add, add_comm] + add_assoc _ _ _ := (toAffineAddEquiv W).injective <| by simp only [map_add, add_assoc] + +end WeierstrassCurve.Projective.Point + namespace WeierstrassCurve.Jacobian.Point /-! ## Weierstrass curves in Jacobian coordinates -/ diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Projective.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Projective.lean index 764adaaef132b..c7c9f0e6a628f 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Projective.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Projective.lean @@ -15,7 +15,8 @@ import Mathlib.Tactic.LinearCombination' This file defines the type of points on a Weierstrass curve as a tuple, consisting of an equivalence class of triples up to scaling by a unit, satisfying a Weierstrass equation with a nonsingular condition. This file also defines the negation and addition operations of the group law for this -type, and proves that they respect the Weierstrass equation and the nonsingular condition. +type, and proves that they respect the Weierstrass equation and the nonsingular condition. The fact +that they form an abelian group is proven in `Mathlib.AlgebraicGeometry.EllipticCurve.Group`. ## Mathematical background From 0a6144da8eda1193d9b71303feefe30c9beb76c3 Mon Sep 17 00:00:00 2001 From: Nailin Guan <150537269+Thmoas-Guan@users.noreply.github.com> Date: Mon, 16 Dec 2024 13:15:11 +0000 Subject: [PATCH 742/829] feat(Topology): Introduce HomeomorphClass (#18689) Introduce HomeomorphClass --- Mathlib/Topology/Algebra/Module/Basic.lean | 4 +++ Mathlib/Topology/Homeomorph.lean | 37 ++++++++++++++++++++++ 2 files changed, 41 insertions(+) diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index 9e0b39d1f56d0..193f14fc095d3 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -309,6 +309,10 @@ instance (priority := 100) continuousSemilinearMapClass [EquivLike F M M₂] [s : ContinuousSemilinearEquivClass F σ M M₂] : ContinuousSemilinearMapClass F σ M M₂ := { s with } +instance (priority := 100) [EquivLike F M M₂] + [s : ContinuousSemilinearEquivClass F σ M M₂] : HomeomorphClass F M M₂ := + { s with } + end ContinuousSemilinearEquivClass section PointwiseLimits diff --git a/Mathlib/Topology/Homeomorph.lean b/Mathlib/Topology/Homeomorph.lean index b7e93c00e5acf..0e4c4e3c97d9c 100644 --- a/Mathlib/Topology/Homeomorph.lean +++ b/Mathlib/Topology/Homeomorph.lean @@ -6,6 +6,7 @@ Authors: Johannes Hölzl, Patrick Massot, Sébastien Gouëzel, Zhouhang Zhou, Re import Mathlib.Logic.Equiv.Fin import Mathlib.Topology.Algebra.Support import Mathlib.Topology.Connected.LocallyConnected +import Mathlib.Topology.ContinuousMap.Defs import Mathlib.Topology.DenseEmbedding /-! @@ -1049,3 +1050,39 @@ lemma IsHomeomorph.pi_map {ι : Type*} {X Y : ι → Type*} [∀ i, TopologicalS [∀ i, TopologicalSpace (Y i)] {f : (i : ι) → X i → Y i} (h : ∀ i, IsHomeomorph (f i)) : IsHomeomorph (fun (x : ∀ i, X i) i ↦ f i (x i)) := (Homeomorph.piCongrRight fun i ↦ (h i).homeomorph (f i)).isHomeomorph + +/-- `HomeomorphClass F A B` states that `F` is a type of homeomorphisms.-/ +class HomeomorphClass (F : Type*) (A B : outParam Type*) + [TopologicalSpace A] [TopologicalSpace B] [h : EquivLike F A B] : Prop where + map_continuous : ∀ (f : F), Continuous f + inv_continuous : ∀ (f : F), Continuous (h.inv f) + +namespace HomeomorphClass + +variable {F α β : Type*} [TopologicalSpace α] [TopologicalSpace β] [EquivLike F α β] + +/-- Turn an element of a type `F` satisfying `HomeomorphClass F α β` into an actual +`Homeomorph`. This is declared as the default coercion from `F` to `α ≃ₜ β`. -/ +@[coe] +def toHomeomorph [h : HomeomorphClass F α β] (f : F) : α ≃ₜ β := + {(f : α ≃ β) with + continuous_toFun := h.map_continuous f + continuous_invFun := h.inv_continuous f } + +@[simp] +theorem coe_coe [h : HomeomorphClass F α β] (f : F) : ⇑(h.toHomeomorph f) = ⇑f := rfl + +instance [HomeomorphClass F α β] : CoeOut F (α ≃ₜ β) := + ⟨HomeomorphClass.toHomeomorph⟩ + +theorem toHomeomorph_injective [HomeomorphClass F α β] : Function.Injective ((↑) : F → α ≃ₜ β) := + fun _ _ e ↦ DFunLike.ext _ _ fun a ↦ congr_arg (fun e : α ≃ₜ β ↦ e.toFun a) e + +instance [HomeomorphClass F α β] : ContinuousMapClass F α β where + map_continuous f := map_continuous f + +instance : HomeomorphClass (α ≃ₜ β) α β where + map_continuous e := e.continuous_toFun + inv_continuous e := e.continuous_invFun + +end HomeomorphClass From 3fc26c460c72cbda03a8327c2bb21befe3d78452 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Mon, 16 Dec 2024 13:15:13 +0000 Subject: [PATCH 743/829] =?UTF-8?q?feat:=20add=20`sq=5Fle=5Fsq=E2=82=80`?= =?UTF-8?q?=20convenience=20lemma=20(#19963)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean | 6 ++++++ Mathlib/Algebra/Order/Ring/Abs.lean | 6 ++---- 2 files changed, 8 insertions(+), 4 deletions(-) diff --git a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean index 01454ee338502..2276d496cf83f 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean @@ -1273,6 +1273,12 @@ lemma lt_of_mul_self_lt_mul_self₀ (hb : 0 ≤ b) : a * a < b * b → a < b := simp_rw [← sq] exact lt_of_pow_lt_pow_left₀ _ hb +lemma sq_lt_sq₀ (ha : 0 ≤ a) (hb : 0 ≤ b) : a ^ 2 < b ^ 2 ↔ a < b := + pow_lt_pow_iff_left₀ ha hb two_ne_zero + +lemma sq_le_sq₀ (ha : 0 ≤ a) (hb : 0 ≤ b) : a ^ 2 ≤ b ^ 2 ↔ a ≤ b := + pow_le_pow_iff_left₀ ha hb two_ne_zero + end MonoidWithZero.LinearOrder section CancelMonoidWithZero diff --git a/Mathlib/Algebra/Order/Ring/Abs.lean b/Mathlib/Algebra/Order/Ring/Abs.lean index 5552705b81dff..68c3cfc6d55b0 100644 --- a/Mathlib/Algebra/Order/Ring/Abs.lean +++ b/Mathlib/Algebra/Order/Ring/Abs.lean @@ -95,15 +95,13 @@ lemma abs_le_one_iff_mul_self_le_one : |a| ≤ 1 ↔ a * a ≤ 1 := by lemma abs_sq (x : α) : |x ^ 2| = x ^ 2 := by simpa only [sq] using abs_mul_self x lemma sq_lt_sq : a ^ 2 < b ^ 2 ↔ |a| < |b| := by - simpa only [sq_abs] using - (pow_left_strictMonoOn₀ two_ne_zero).lt_iff_lt (abs_nonneg a) (abs_nonneg b) + simpa only [sq_abs] using sq_lt_sq₀ (abs_nonneg a) (abs_nonneg b) lemma sq_lt_sq' (h1 : -b < a) (h2 : a < b) : a ^ 2 < b ^ 2 := sq_lt_sq.2 (lt_of_lt_of_le (abs_lt.2 ⟨h1, h2⟩) (le_abs_self _)) lemma sq_le_sq : a ^ 2 ≤ b ^ 2 ↔ |a| ≤ |b| := by - simpa only [sq_abs] using - (pow_left_strictMonoOn₀ two_ne_zero).le_iff_le (abs_nonneg a) (abs_nonneg b) + simpa only [sq_abs] using sq_le_sq₀ (abs_nonneg a) (abs_nonneg b) lemma sq_le_sq' (h1 : -b ≤ a) (h2 : a ≤ b) : a ^ 2 ≤ b ^ 2 := sq_le_sq.2 (le_trans (abs_le.mpr ⟨h1, h2⟩) (le_abs_self _)) From 4e1ec44b6955eaf06bfeea70d13a04212ba9c442 Mon Sep 17 00:00:00 2001 From: smorel394 Date: Mon, 16 Dec 2024 13:49:38 +0000 Subject: [PATCH 744/829] feat(Algebra/Category/Grp/Limits): the forgetful functor on the category of groups creates all limits (#19966) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The forgetful functor on the category of groups (commutative or not) creates all limits. As we already know that the forgetful functor from groups to monoids creates all limits (once we remove the unnecessary size condition in `Grp.Forget₂.createsLimit`), this just follows from the similar fact for monoid categories, which is PR #19965. - [x] depends on: #19965 Co-authored-by: Joël Riou Co-authored-by: smorel394 <67864981+smorel394@users.noreply.github.com> --- Mathlib/Algebra/Category/Grp/Limits.lean | 69 +++++++++++++++++++----- 1 file changed, 55 insertions(+), 14 deletions(-) diff --git a/Mathlib/Algebra/Category/Grp/Limits.lean b/Mathlib/Algebra/Category/Grp/Limits.lean index f9226f51a1742..f5aac619ae4f9 100644 --- a/Mathlib/Algebra/Category/Grp/Limits.lean +++ b/Mathlib/Algebra/Category/Grp/Limits.lean @@ -3,13 +3,11 @@ Copyright (c) 2020 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ -import Mathlib.Algebra.Category.MonCat.Limits import Mathlib.Algebra.Category.Grp.ForgetCorepresentable import Mathlib.Algebra.Category.Grp.Preadditive -import Mathlib.Algebra.Group.Subgroup.Basic -import Mathlib.CategoryTheory.Comma.Over -import Mathlib.CategoryTheory.Limits.ConcreteCategory.Basic +import Mathlib.Algebra.Category.MonCat.Limits import Mathlib.CategoryTheory.ConcreteCategory.ReflectsIso +import Mathlib.CategoryTheory.Limits.ConcreteCategory.Basic /-! # The category of (commutative) (additive) groups has all limits @@ -87,16 +85,21 @@ noncomputable instance Forget₂.createsLimit : CategoryTheory.reflectsIsomorphisms_forget₂ _ _ createsLimitOfReflectsIso (K := F) (F := (forget₂ Grp.{u} MonCat.{u})) fun c' t => - { liftedCone := - { pt := Grp.of (Types.Small.limitCone (F ⋙ forget Grp)).pt - π := - { app := MonCat.limitπMonoidHom (F ⋙ forget₂ Grp MonCat) - naturality := - (MonCat.HasLimits.limitCone - (F ⋙ forget₂ Grp MonCat.{u})).π.naturality } } - validLift := by apply IsLimit.uniqueUpToIso (MonCat.HasLimits.limitConeIsLimit.{v, u} _) t - makesLimit := - IsLimit.ofFaithful (forget₂ Grp MonCat.{u}) (MonCat.HasLimits.limitConeIsLimit _) + have : Small.{u} (Functor.sections ((F ⋙ forget₂ Grp MonCat) ⋙ forget MonCat)) := by + have : HasLimit (F ⋙ forget₂ Grp MonCat) := ⟨_, t⟩ + apply Concrete.small_sections_of_hasLimit (F ⋙ forget₂ Grp MonCat) + have : Small.{u} (Functor.sections (F ⋙ forget Grp)) := inferInstanceAs <| Small.{u} + (Functor.sections ((F ⋙ forget₂ Grp MonCat) ⋙ forget MonCat)) + { liftedCone := + { pt := Grp.of (Types.Small.limitCone (F ⋙ forget Grp)).pt + π := + { app := MonCat.limitπMonoidHom (F ⋙ forget₂ Grp MonCat) + naturality := + (MonCat.HasLimits.limitCone + (F ⋙ forget₂ Grp MonCat.{u})).π.naturality } } + validLift := by apply IsLimit.uniqueUpToIso (MonCat.HasLimits.limitConeIsLimit.{v, u} _) t + makesLimit := + IsLimit.ofFaithful (forget₂ Grp MonCat.{u}) (MonCat.HasLimits.limitConeIsLimit _) (fun _ => _) fun _ => rfl } /-- A choice of limit cone for a functor into `Grp`. @@ -193,6 +196,25 @@ instance forget_preservesLimitsOfSize : instance forget_preservesLimits : PreservesLimits (forget Grp.{u}) := Grp.forget_preservesLimitsOfSize.{u, u} +@[to_additive] +noncomputable instance forget_createsLimit : + CreatesLimit F (forget Grp.{u}) := by + set e : forget₂ Grp.{u} MonCat.{u} ⋙ forget MonCat.{u} ≅ forget Grp.{u} := Iso.refl _ + exact createsLimitOfNatIso e + +@[to_additive] +noncomputable instance forget_createsLimitsOfShape : + CreatesLimitsOfShape J (forget Grp.{u}) where + CreatesLimit := inferInstance + +/-- The forgetful functor from groups to types creates all limits. +-/ +@[to_additive + "The forgetful functor from additive groups to types creates all limits.", + to_additive_relevant_arg 2] +noncomputable instance forget_createsLimitsOfSize : + CreatesLimitsOfSize.{w, v} (forget Grp.{u}) where + CreatesLimitsOfShape := inferInstance end Grp namespace CommGrp @@ -398,6 +420,25 @@ noncomputable instance _root_.AddCommGrp.forget_preservesLimits : noncomputable instance forget_preservesLimits : PreservesLimits (forget CommGrp.{u}) := CommGrp.forget_preservesLimitsOfSize.{u, u} +@[to_additive] +noncomputable instance forget_createsLimit : + CreatesLimit F (forget CommGrp.{u}) := by + set e : forget₂ CommGrp.{u} Grp.{u} ⋙ forget Grp.{u} ≅ forget CommGrp.{u} := Iso.refl _ + exact createsLimitOfNatIso e + +@[to_additive] +noncomputable instance forget_createsLimitsOfShape (J : Type v) [Category.{w} J] : + CreatesLimitsOfShape J (forget CommGrp.{u}) where + CreatesLimit := inferInstance + +/-- The forgetful functor from commutative groups to types creates all limits. +-/ +@[to_additive + "The forgetful functor from additive commutative groups to types creates all limits.", + to_additive_relevant_arg 2] +noncomputable instance forget_createsLimitsOfSize : + CreatesLimitsOfSize.{w, v} (forget CommGrp.{u}) where + CreatesLimitsOfShape := inferInstance -- Verify we can form limits indexed over smaller categories. example (f : ℕ → AddCommGrp) : HasProduct f := by infer_instance From 4c066650b87bc861c47ce5b4d83bcedf972e22eb Mon Sep 17 00:00:00 2001 From: David Renshaw Date: Mon, 16 Dec 2024 14:15:35 +0000 Subject: [PATCH 745/829] chore(RingTheory/Localization): prove isUnit_den_iff.mp by isInteger_of_isUnit_den (#19923) Replaces the 9-line proof of `isUnit_den_iff.mp` with a direct application of the existing lemma `sInteger_of_isUnit_den` (which is in fact the preceding item in the file). For reference, `isUnit_den_iff` was first added in #14643. This simplification was found by [`tryAtEachStep`](https://github.com/dwrensha/tryAtEachStep). --- Mathlib/RingTheory/Localization/NumDen.lean | 11 +---------- 1 file changed, 1 insertion(+), 10 deletions(-) diff --git a/Mathlib/RingTheory/Localization/NumDen.lean b/Mathlib/RingTheory/Localization/NumDen.lean index 290ca0e000441..1df3ce301fd17 100644 --- a/Mathlib/RingTheory/Localization/NumDen.lean +++ b/Mathlib/RingTheory/Localization/NumDen.lean @@ -105,16 +105,7 @@ theorem isInteger_of_isUnit_den {x : K} (h : IsUnit (den A x : A)) : IsInteger A rw [← mul_assoc, mul_inv_cancel₀ d_ne_zero, one_mul, mk'_spec'] theorem isUnit_den_iff (x : K) : IsUnit (den A x : A) ↔ IsLocalization.IsInteger A x where - mp h := by - obtain ⟨den, hd⟩ := IsUnit.exists_right_inv h - use (num A x) * den - conv => rhs; rw [← mk'_num_den' A x] - rw [map_mul, div_eq_mul_inv] - congr 1 - apply eq_inv_of_mul_eq_one_right - rw [← map_mul] - norm_cast - simp only [hd, OneMemClass.coe_one, map_one] + mp := isInteger_of_isUnit_den mpr h := by have ⟨v, h⟩ := h apply IsRelPrime.isUnit_of_dvd (num_den_reduced A x).symm From 2927e227ec9454096c0507b4ac8386b3f4f09e5e Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Mon, 16 Dec 2024 14:33:32 +0000 Subject: [PATCH 746/829] feat(LinearAlgebra): general transitivity of norm (#19978) **RingTheory/AlgebraTower.lean**: add a version of `Basis.smulTower` with the two index types swapped and associated lemmas; generalize from Ring/Group to Semiring/Monoid. **LinearAlgebra/Matrix/ToLin.lean**: add a lemma on the matrix representation of `LinearMap.restrictScalars` w.r.t. `Basis.smulTower'`; generalize from Ring/Group to Semiring/Monoid. **RingTheory/Norm/Transitivity.lean**: add the result in Silvester's paper **Determinants of block matrices** (with commuting, equal-sized, square blocks). The determinant can be computed by taking two determinants in a row: first over the commutative ring generated by the blocks, then over the base ring. As consequences, the determinant of `LinearMap.restrictScalars` can be computed by first taking a determinant and then taking a norm, and the norm satisfies transitivity in a tower of algebras. **RingTheory/Norm/Basic.lean**: `Algebra.norm_norm` is replaced by the more general version proven in **Transitivity.lean**. **Algebra/BigOperators/Group/Finset.lean**, **Data/Matrix/Basic.lean**, **LinearAlgebra/Determinant.lean**, **RingTheory/Finiteness/Cardinality.lean**: add one lemma in each. **Data/Matrix/Block.lean**: add 3 lemmas. **LinearAlgebra/Matrix/Block.lean**: generalize `BlockTriangular` from `CommRing` to `Zero` and add two lemmas. **LinearAlgebra/Matrix/Determinant/Basic.lean**: somehow a proof breaks and has to be fixed. Co-authored-by: Eric Wieser --- Mathlib.lean | 1 + .../Algebra/BigOperators/Group/Finset.lean | 3 + Mathlib/Data/Matrix/Block.lean | 29 ++- Mathlib/Data/Matrix/Composition.lean | 20 +- Mathlib/Data/Matrix/Defs.lean | 10 + Mathlib/LinearAlgebra/Determinant.lean | 4 + Mathlib/LinearAlgebra/Matrix/Block.lean | 59 ++++-- .../Matrix/Determinant/Basic.lean | 8 +- Mathlib/LinearAlgebra/Matrix/ToLin.lean | 20 +- Mathlib/NumberTheory/NumberField/Norm.lean | 1 + Mathlib/RingTheory/AlgebraTower.lean | 35 +++- .../RingTheory/Finiteness/Cardinality.lean | 6 + Mathlib/RingTheory/Norm/Basic.lean | 38 ---- Mathlib/RingTheory/Norm/Transitivity.lean | 196 ++++++++++++++++++ Mathlib/RingTheory/PolynomialAlgebra.lean | 17 ++ docs/references.bib | 13 ++ 16 files changed, 382 insertions(+), 78 deletions(-) create mode 100644 Mathlib/RingTheory/Norm/Transitivity.lean diff --git a/Mathlib.lean b/Mathlib.lean index 87d14e789d6f0..00329c87c9e80 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4499,6 +4499,7 @@ import Mathlib.RingTheory.NonUnitalSubsemiring.Basic import Mathlib.RingTheory.NonUnitalSubsemiring.Defs import Mathlib.RingTheory.Norm.Basic import Mathlib.RingTheory.Norm.Defs +import Mathlib.RingTheory.Norm.Transitivity import Mathlib.RingTheory.NormTrace import Mathlib.RingTheory.Nullstellensatz import Mathlib.RingTheory.OreLocalization.Basic diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index c31cf3896b7ed..80ea8b0b9c9f1 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -1995,6 +1995,9 @@ theorem prod_subsingleton {α β : Type*} [CommMonoid β] [Subsingleton α] [Fin have : Unique α := uniqueOfSubsingleton a rw [prod_unique f, Subsingleton.elim default a] +@[to_additive] theorem prod_Prop {β} [CommMonoid β] (f : Prop → β) : + ∏ p, f p = f True * f False := by simp + @[to_additive] theorem prod_subtype_mul_prod_subtype {α β : Type*} [Fintype α] [CommMonoid β] (p : α → Prop) (f : α → β) [DecidablePred p] : diff --git a/Mathlib/Data/Matrix/Block.lean b/Mathlib/Data/Matrix/Block.lean index cb4875d44825f..15185b1ef2dbb 100644 --- a/Mathlib/Data/Matrix/Block.lean +++ b/Mathlib/Data/Matrix/Block.lean @@ -3,6 +3,8 @@ Copyright (c) 2018 Ellen Arlt. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Ellen Arlt, Blair Shi, Sean Leather, Mario Carneiro, Johan Commelin -/ +import Mathlib.Data.Matrix.Basic +import Mathlib.Data.Matrix.Composition import Mathlib.Data.Matrix.ConjTranspose /-! @@ -21,7 +23,6 @@ import Mathlib.Data.Matrix.ConjTranspose * `Matrix.blockDiag'`: extract the blocks from the diagonal of a block diagonal matrix. -/ - variable {l m n o p q : Type*} {m' n' p' : o → Type*} variable {R : Type*} {S : Type*} {α : Type*} {β : Type*} @@ -833,3 +834,29 @@ theorem toBlock_mul_eq_add {m n k : Type*} [Fintype n] (p : m → Prop) (q : n end end Matrix + +section Maps + +variable {R α β ι : Type*} + +lemma Matrix.map_toSquareBlock + (f : α → β) {M : Matrix m m α} {ι} {b : m → ι} {i : ι} : + (M.map f).toSquareBlock b i = (M.toSquareBlock b i).map f := + submatrix_map _ _ _ _ + +lemma Matrix.comp_toSquareBlock {b : m → α} + (M : Matrix m m (Matrix n n R)) (a : α) : + letI equiv := Equiv.prodSubtypeFstEquivSubtypeProd.symm + (M.comp m m n n R).toSquareBlock (fun i ↦ b i.1) a = + ((M.toSquareBlock b a).comp _ _ n n R).reindex equiv equiv := + rfl + +variable [Zero R] [DecidableEq m] + +lemma Matrix.comp_diagonal (d) : + comp m m n n R (diagonal d) = + (blockDiagonal d).reindex (.prodComm ..) (.prodComm ..) := by + ext + simp [diagonal, blockDiagonal, Matrix.ite_apply] + +end Maps diff --git a/Mathlib/Data/Matrix/Composition.lean b/Mathlib/Data/Matrix/Composition.lean index 1970544baa99b..b1d7909c4817e 100644 --- a/Mathlib/Data/Matrix/Composition.lean +++ b/Mathlib/Data/Matrix/Composition.lean @@ -40,23 +40,37 @@ section AddCommMonoid variable [AddCommMonoid R] /-- `Matrix.comp` as `AddEquiv` -/ -@[simps!] def compAddEquiv : Matrix I J (Matrix K L R) ≃+ Matrix (I × K) (J × L) R where __ := Matrix.comp I J K L R map_add' _ _ := rfl +@[simp] +theorem compAddEquiv_apply (M : Matrix I J (Matrix K L R)) : + compAddEquiv I J K L R M = comp I J K L R M := rfl + +@[simp] +theorem compAddEquiv_symm_apply (M : Matrix (I × K) (J × L) R) : + (compAddEquiv I J K L R).symm M = (comp I J K L R).symm M := rfl + end AddCommMonoid section Semiring -variable [Semiring R] [Fintype I] [Fintype J] [DecidableEq I] [DecidableEq J] +variable [Semiring R] [Fintype I] [Fintype J] /-- `Matrix.comp` as `RingEquiv` -/ -@[simps!] def compRingEquiv : Matrix I I (Matrix J J R) ≃+* Matrix (I × J) (I × J) R where __ := Matrix.compAddEquiv I I J J R map_mul' _ _ := by ext; exact (Matrix.sum_apply ..).trans <| .symm <| Fintype.sum_prod_type .. +@[simp] +theorem compRingEquiv_apply (M : Matrix I I (Matrix J J R)) : + compRingEquiv I J R M = comp I I J J R M := rfl + +@[simp] +theorem compRingEquiv_symm_apply (M : Matrix (I × J) (I × J) R) : + (compRingEquiv I J R).symm M = (comp I I J J R).symm M := rfl + end Semiring section LinearMap diff --git a/Mathlib/Data/Matrix/Defs.lean b/Mathlib/Data/Matrix/Defs.lean index 7767c0b0b6d61..d85e8a18a4b44 100644 --- a/Mathlib/Data/Matrix/Defs.lean +++ b/Mathlib/Data/Matrix/Defs.lean @@ -237,6 +237,16 @@ theorem sub_apply [Sub α] (A B : Matrix m n α) (i : m) (j : n) : theorem neg_apply [Neg α] (A : Matrix m n α) (i : m) (j : n) : (-A) i j = -(A i j) := rfl +protected theorem dite_apply (P : Prop) [Decidable P] + (A : P → Matrix m n α) (B : ¬P → Matrix m n α) (i : m) (j : n) : + dite P A B i j = dite P (A · i j) (B · i j) := by + rw [dite_apply, dite_apply] + +protected theorem ite_apply (P : Prop) [Decidable P] + (A : Matrix m n α) (B : Matrix m n α) (i : m) (j : n) : + (if P then A else B) i j = if P then A i j else B i j := + Matrix.dite_apply _ _ _ _ _ + end /-! simp-normal form pulls `of` to the outside. -/ diff --git a/Mathlib/LinearAlgebra/Determinant.lean b/Mathlib/LinearAlgebra/Determinant.lean index f1f5ef6f34822..6a4b2827498d0 100644 --- a/Mathlib/LinearAlgebra/Determinant.lean +++ b/Mathlib/LinearAlgebra/Determinant.lean @@ -257,6 +257,10 @@ theorem det_zero [Module.Free A M] : LinearMap.det (0 : M →ₗ[A] M) = (0 : A) ^ Module.finrank A M := by simp only [← zero_smul A (1 : M →ₗ[A] M), det_smul, mul_one, MonoidHom.map_one] +theorem det_eq_one_of_not_module_finite (h : ¬Module.Finite R M) (f : M →ₗ[R] M) : f.det = 1 := by + rw [LinearMap.det, dif_neg, MonoidHom.one_apply] + exact fun ⟨_, ⟨b⟩⟩ ↦ h (Module.Finite.of_basis b) + theorem det_eq_one_of_subsingleton [Subsingleton M] (f : M →ₗ[R] M) : LinearMap.det (f : M →ₗ[R] M) = 1 := by have b : Basis (Fin 0) R M := Basis.empty M diff --git a/Mathlib/LinearAlgebra/Matrix/Block.lean b/Mathlib/LinearAlgebra/Matrix/Block.lean index 26de8f3e90de3..6a494a412ea2e 100644 --- a/Mathlib/LinearAlgebra/Matrix/Block.lean +++ b/Mathlib/LinearAlgebra/Matrix/Block.lean @@ -39,7 +39,7 @@ open Matrix universe v variable {α β m n o : Type*} {m' n' : α → Type*} -variable {R : Type v} [CommRing R] {M N : Matrix m m R} {b : m → α} +variable {R : Type v} {M N : Matrix m m R} {b : m → α} namespace Matrix @@ -47,6 +47,10 @@ section LT variable [LT α] +section Zero + +variable [Zero R] + /-- Let `b` map rows and columns of a square matrix `M` to blocks indexed by `α`s. Then `BlockTriangular M n b` says the matrix is block triangular. -/ def BlockTriangular (M : Matrix m m R) (b : m → α) : Prop := @@ -76,33 +80,49 @@ protected theorem blockTriangular_transpose_iff {b : m → αᵒᵈ} : @[simp] theorem blockTriangular_zero : BlockTriangular (0 : Matrix m m R) b := fun _ _ _ => rfl -protected theorem BlockTriangular.neg (hM : BlockTriangular M b) : BlockTriangular (-M) b := - fun _ _ h => neg_eq_zero.2 <| hM h +end Zero + +protected theorem BlockTriangular.neg [NegZeroClass R] {M : Matrix m m R} + (hM : BlockTriangular M b) : BlockTriangular (-M) b := + fun _ _ h => by rw [neg_apply, hM h, neg_zero] -theorem BlockTriangular.add (hM : BlockTriangular M b) (hN : BlockTriangular N b) : +theorem BlockTriangular.add [AddZeroClass R] (hM : BlockTriangular M b) (hN : BlockTriangular N b) : BlockTriangular (M + N) b := fun i j h => by simp_rw [Matrix.add_apply, hM h, hN h, zero_add] -theorem BlockTriangular.sub (hM : BlockTriangular M b) (hN : BlockTriangular N b) : +theorem BlockTriangular.sub [SubNegZeroMonoid R] + (hM : BlockTriangular M b) (hN : BlockTriangular N b) : BlockTriangular (M - N) b := fun i j h => by simp_rw [Matrix.sub_apply, hM h, hN h, sub_zero] -lemma BlockTriangular.add_iff_right (hM : BlockTriangular M b) : - BlockTriangular (M + N) b ↔ BlockTriangular N b := ⟨(by simpa using ·.sub hM), hM.add⟩ +lemma BlockTriangular.add_iff_right [AddGroup R] (hM : BlockTriangular M b) : + BlockTriangular (M + N) b ↔ BlockTriangular N b := ⟨(by simpa using hM.neg.add ·), hM.add⟩ -lemma BlockTriangular.add_iff_left (hN : BlockTriangular N b) : - BlockTriangular (M + N) b ↔ BlockTriangular M b := by rw [add_comm, hN.add_iff_right] +lemma BlockTriangular.add_iff_left [AddGroup R] (hN : BlockTriangular N b) : + BlockTriangular (M + N) b ↔ BlockTriangular M b := ⟨(by simpa using ·.sub hN), (·.add hN)⟩ -lemma BlockTriangular.sub_iff_right (hM : BlockTriangular M b) : - BlockTriangular (M - N) b ↔ BlockTriangular N b := ⟨(by simpa using hM.sub ·), hM.sub⟩ +lemma BlockTriangular.sub_iff_right [AddGroup R] (hM : BlockTriangular M b) : + BlockTriangular (M - N) b ↔ BlockTriangular N b := ⟨(by simpa using ·.neg.add hM), hM.sub⟩ -lemma BlockTriangular.sub_iff_left (hN : BlockTriangular N b) : +lemma BlockTriangular.sub_iff_left [AddGroup R] (hN : BlockTriangular N b) : BlockTriangular (M - N) b ↔ BlockTriangular M b := ⟨(by simpa using ·.add hN), (·.sub hN)⟩ +lemma BlockTriangular.map {S F} [FunLike F R S] [Zero R] [Zero S] [ZeroHomClass F R S] (f : F) + (h : BlockTriangular M b) : BlockTriangular (M.map f) b := + fun i j lt ↦ by simp [h lt] + +lemma BlockTriangular.comp [Zero R] {M : Matrix m m (Matrix n n R)} (h : BlockTriangular M b) : + BlockTriangular (M.comp m m n n R) fun i ↦ b i.1 := + fun i j lt ↦ by simp [h lt] + end LT section Preorder variable [Preorder α] +section Zero + +variable [Zero R] + theorem blockTriangular_diagonal [DecidableEq m] (d : m → R) : BlockTriangular (diagonal d) b := fun _ _ h => diagonal_apply_ne' d fun h' => ne_of_lt h (congr_arg _ h') @@ -119,7 +139,7 @@ theorem blockTriangular_blockDiagonal [DecidableEq α] (d : α → Matrix m m R) variable [DecidableEq m] -theorem blockTriangular_one : BlockTriangular (1 : Matrix m m R) b := +theorem blockTriangular_one [One R] : BlockTriangular (1 : Matrix m m R) b := blockTriangular_diagonal _ theorem blockTriangular_stdBasisMatrix {i j : m} (hij : b i ≤ b j) (c : R) : @@ -133,6 +153,10 @@ theorem blockTriangular_stdBasisMatrix' {i j : m} (hij : b j ≤ b i) (c : R) : BlockTriangular (stdBasisMatrix i j c) (toDual ∘ b) := blockTriangular_stdBasisMatrix (by exact toDual_le_toDual.mpr hij) _ +end Zero + +variable [CommRing R] [DecidableEq m] + theorem blockTriangular_transvection {i j : m} (hij : b i ≤ b j) (c : R) : BlockTriangular (transvection i j c) b := blockTriangular_one.add (blockTriangular_stdBasisMatrix hij c) @@ -147,7 +171,8 @@ section LinearOrder variable [LinearOrder α] -theorem BlockTriangular.mul [Fintype m] {M N : Matrix m m R} (hM : BlockTriangular M b) +theorem BlockTriangular.mul [Fintype m] [NonUnitalNonAssocSemiring R] + {M N : Matrix m m R} (hM : BlockTriangular M b) (hN : BlockTriangular N b) : BlockTriangular (M * N) b := by intro i j hij apply Finset.sum_eq_zero @@ -158,7 +183,7 @@ theorem BlockTriangular.mul [Fintype m] {M N : Matrix m m R} (hM : BlockTriangul end LinearOrder -theorem upper_two_blockTriangular [Preorder α] (A : Matrix m m R) (B : Matrix m n R) +theorem upper_two_blockTriangular [Zero R] [Preorder α] (A : Matrix m m R) (B : Matrix m n R) (D : Matrix n n R) {a b : α} (hab : a < b) : BlockTriangular (fromBlocks A B 0 D) (Sum.elim (fun _ => a) fun _ => b) := by rintro (c | c) (d | d) hcd <;> first | simp [hab.not_lt] at hcd ⊢ @@ -166,7 +191,7 @@ theorem upper_two_blockTriangular [Preorder α] (A : Matrix m m R) (B : Matrix m /-! ### Determinant -/ -variable [DecidableEq m] [Fintype m] [DecidableEq n] [Fintype n] +variable [CommRing R] [DecidableEq m] [Fintype m] [DecidableEq n] [Fintype n] theorem equiv_block_det (M : Matrix m m R) {p q : m → Prop} [DecidablePred p] [DecidablePred q] (e : ∀ x, q x ↔ p x) : (toSquareBlockProp M p).det = (toSquareBlockProp M q).det := by @@ -257,7 +282,7 @@ theorem det_of_lowerTriangular [LinearOrder m] (M : Matrix m m R) (h : M.BlockTr open Polynomial -theorem matrixOfPolynomials_blockTriangular {n : ℕ} (p : Fin n → R[X]) +theorem matrixOfPolynomials_blockTriangular {R} [Semiring R] {n : ℕ} (p : Fin n → R[X]) (h_deg : ∀ i, (p i).natDegree ≤ i) : Matrix.BlockTriangular (Matrix.of (fun (i j : Fin n) => (p j).coeff i)) id := fun _ j h => by diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean index 274f69211c711..544a0f20ffac2 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean @@ -790,10 +790,10 @@ theorem det_fin_three (A : Matrix (Fin 3) (Fin 3) R) : A 0 0 * A 1 1 * A 2 2 - A 0 0 * A 1 2 * A 2 1 - A 0 1 * A 1 0 * A 2 2 + A 0 1 * A 1 2 * A 2 0 + A 0 2 * A 1 0 * A 2 1 - A 0 2 * A 1 1 * A 2 0 := by - 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.val_eq_zero, Fin.succ_succAbove_zero, sum_singleton, Fin.succ_succAbove_one, even_add_self] + simp only [det_succ_row_zero, submatrix_apply, Fin.succ_zero_eq_one, submatrix_submatrix, + det_unique, Fin.default_eq_zero, Function.comp_apply, Fin.succ_one_eq_two, Fin.sum_univ_succ, + Fin.val_zero, Fin.zero_succAbove, univ_unique, Fin.val_succ, Fin.val_eq_zero, + Fin.succ_succAbove_zero, sum_singleton, Fin.succ_succAbove_one] ring end Matrix diff --git a/Mathlib/LinearAlgebra/Matrix/ToLin.lean b/Mathlib/LinearAlgebra/Matrix/ToLin.lean index de85599206032..1ef1d74c925eb 100644 --- a/Mathlib/LinearAlgebra/Matrix/ToLin.lean +++ b/Mathlib/LinearAlgebra/Matrix/ToLin.lean @@ -794,7 +794,7 @@ namespace Algebra section Lmul -variable {R S : Type*} [CommRing R] [Ring S] [Algebra R S] +variable {R S : Type*} [CommSemiring R] [Semiring S] [Algebra R S] variable {m : Type*} [Fintype m] [DecidableEq m] (b : Basis m R S) theorem toMatrix_lmul' (x : S) (i j) : @@ -865,11 +865,21 @@ theorem smul_leftMulMatrix {G} [Group G] [DistribMulAction G S] Basis.repr_smul, Basis.smul_apply, LinearEquiv.trans_apply, DistribMulAction.toLinearEquiv_symm_apply, mul_smul_comm, inv_smul_smul] +variable {A M n : Type*} [Fintype n] [DecidableEq n] + [CommSemiring A] [AddCommMonoid M] [Module R M] [Module A M] [Algebra R A] [IsScalarTower R A M] + (bA : Basis m R A) (bM : Basis n A M) + +lemma _root_.LinearMap.restrictScalars_toMatrix (f : M →ₗ[A] M) : + (f.restrictScalars R).toMatrix (bA.smulTower' bM) (bA.smulTower' bM) = + ((f.toMatrix bM bM).map (leftMulMatrix bA)).comp _ _ _ _ _ := by + ext; simp [toMatrix, Basis.repr, Algebra.leftMulMatrix_apply, + Basis.smulTower'_repr, Basis.smulTower'_apply, mul_comm] + end Lmul section LmulTower -variable {R S T : Type*} [CommRing R] [CommRing S] [Ring T] +variable {R S T : Type*} [CommSemiring R] [CommSemiring S] [Semiring T] variable [Algebra R S] [Algebra S T] [Algebra R T] [IsScalarTower R S T] variable {m n : Type*} [Fintype m] [Fintype n] [DecidableEq m] [DecidableEq n] variable (b : Basis m R S) (c : Basis n S T) @@ -901,9 +911,9 @@ end Algebra section -variable {R : Type*} [CommRing R] {n : Type*} [DecidableEq n] -variable {M M₁ M₂ : Type*} [AddCommGroup M] [Module R M] -variable [AddCommGroup M₁] [Module R M₁] [AddCommGroup M₂] [Module R M₂] +variable {R : Type*} [CommSemiring R] {n : Type*} [DecidableEq n] +variable {M M₁ M₂ : Type*} [AddCommMonoid M] [Module R M] +variable [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] /-- The natural equivalence between linear endomorphisms of finite free modules and square matrices is compatible with the algebra structures. -/ diff --git a/Mathlib/NumberTheory/NumberField/Norm.lean b/Mathlib/NumberTheory/NumberField/Norm.lean index e8964a687a66b..5db370bf1a755 100644 --- a/Mathlib/NumberTheory/NumberField/Norm.lean +++ b/Mathlib/NumberTheory/NumberField/Norm.lean @@ -5,6 +5,7 @@ Authors: Riccardo Brasca, Eric Rodriguez -/ import Mathlib.NumberTheory.NumberField.Basic import Mathlib.RingTheory.Localization.NormTrace +import Mathlib.RingTheory.Norm.Transitivity /-! # Norm in number fields diff --git a/Mathlib/RingTheory/AlgebraTower.lean b/Mathlib/RingTheory/AlgebraTower.lean index e08bb34ab5977..abfe65fb0a6fe 100644 --- a/Mathlib/RingTheory/AlgebraTower.lean +++ b/Mathlib/RingTheory/AlgebraTower.lean @@ -94,6 +94,8 @@ open scoped Classical universe v₁ w₁ variable {R S A} + +section variable [Ring R] [Ring S] [AddCommGroup A] variable [Module R S] [Module S A] [Module R A] [IsScalarTower R S A] @@ -111,8 +113,11 @@ theorem linearIndependent_smul {ι : Type v₁} {b : ι → S} {ι' : Type w₁} simp_rw [← smul_assoc, ← Finset.sum_smul] at h1 exact hb _ _ (hc _ _ h1 k (Finset.mem_image_of_mem _ hik)) i (Finset.mem_image_of_mem _ hik) exact hg _ hik +end variable (R) +variable [Semiring R] [Semiring S] [AddCommMonoid A] +variable [Module R S] [Module S A] [Module R A] [IsScalarTower R S A] -- LinearIndependent is enough if S is a ring rather than semiring. theorem Basis.isScalarTower_of_nonempty {ι} [Nonempty ι] (b : Basis ι S A) : IsScalarTower R S S := @@ -122,13 +127,12 @@ theorem Basis.isScalarTower_of_nonempty {ι} [Nonempty ι] (b : Basis ι S A) : theorem Basis.isScalarTower_finsupp {ι} (b : Basis ι S A) : IsScalarTower R S (ι →₀ S) := b.repr.symm.isScalarTower_of_injective R b.repr.symm.injective -variable {R} +variable {R} {ι ι' : Type*} (b : Basis ι R S) (c : Basis ι' S A) /-- `Basis.smulTower (b : Basis ι R S) (c : Basis ι S A)` is the `R`-basis on `A` where the `(i, j)`th basis vector is `b i • c j`. -/ noncomputable -def Basis.smulTower {ι : Type v₁} {ι' : Type w₁} (b : Basis ι R S) (c : Basis ι' S A) : - Basis (ι × ι') R A := +def Basis.smulTower : Basis (ι × ι') R A := haveI := c.isScalarTower_finsupp R .ofRepr (c.repr.restrictScalars R ≪≫ₗ @@ -137,19 +141,15 @@ def Basis.smulTower {ι : Type v₁} {ι' : Type w₁} (b : Basis ι R S) (c : B Finsupp.lcongr (Equiv.prodComm ι' ι) (LinearEquiv.refl _ _)))) @[simp] -theorem Basis.smulTower_repr {ι : Type v₁} {ι' : Type w₁} - (b : Basis ι R S) (c : Basis ι' S A) (x ij) : +theorem Basis.smulTower_repr (x ij) : (b.smulTower c).repr x ij = b.repr (c.repr x ij.2) ij.1 := by simp [smulTower] -theorem Basis.smulTower_repr_mk {ι : Type v₁} {ι' : Type w₁} (b : Basis ι R S) (c : Basis ι' S A) - (x i j) : (b.smulTower c).repr x (i, j) = b.repr (c.repr x j) i := +theorem Basis.smulTower_repr_mk (x i j) : (b.smulTower c).repr x (i, j) = b.repr (c.repr x j) i := b.smulTower_repr c x (i, j) @[simp] -theorem Basis.smulTower_apply {ι : Type v₁} {ι' : Type w₁} - (b : Basis ι R S) (c : Basis ι' S A) (ij) : - (b.smulTower c) ij = b ij.1 • c ij.2 := by +theorem Basis.smulTower_apply (ij) : (b.smulTower c) ij = b ij.1 • c ij.2 := by obtain ⟨i, j⟩ := ij rw [Basis.apply_eq_iff] ext ⟨i', j'⟩ @@ -160,6 +160,21 @@ theorem Basis.smulTower_apply {ι : Type v₁} {ι' : Type w₁} · simp [hi, Finsupp.single_apply] · simp [hi] +/-- `Basis.smulTower (b : Basis ι R S) (c : Basis ι S A)` is the `R`-basis on `A` +where the `(i, j)`th basis vector is `b j • c i`. -/ +noncomputable def Basis.smulTower' : Basis (ι' × ι) R A := + (b.smulTower c).reindex (.prodComm ..) + +theorem Basis.smulTower'_repr (x ij) : + (b.smulTower' c).repr x ij = b.repr (c.repr x ij.1) ij.2 := by + rw [smulTower', repr_reindex_apply, smulTower_repr]; rfl + +theorem Basis.smulTower'_repr_mk (x i j) : (b.smulTower' c).repr x (i, j) = b.repr (c.repr x i) j := + b.smulTower'_repr c x (i, j) + +theorem Basis.smulTower'_apply (ij) : b.smulTower' c ij = b ij.2 • c ij.1 := by + rw [smulTower', reindex_apply, smulTower_apply]; rfl + end Semiring section Ring diff --git a/Mathlib/RingTheory/Finiteness/Cardinality.lean b/Mathlib/RingTheory/Finiteness/Cardinality.lean index 630c9f5a5c419..e03f6de91613f 100644 --- a/Mathlib/RingTheory/Finiteness/Cardinality.lean +++ b/Mathlib/RingTheory/Finiteness/Cardinality.lean @@ -67,8 +67,14 @@ lemma finite_basis [Nontrivial R] {ι} [Module.Finite R M] _root_.Finite ι := let ⟨s, hs⟩ := ‹Module.Finite R M› basis_finite_of_finite_spans (↑s) s.finite_toSet hs b + end Finite +variable {R M} +lemma not_finite_of_infinite_basis [Nontrivial R] {ι} [Infinite ι] (b : Basis ι R M) : + ¬ Module.Finite R M := + fun _ ↦ (Finite.finite_basis b).not_infinite ‹_› + end Module end ModuleAndAlgebra diff --git a/Mathlib/RingTheory/Norm/Basic.lean b/Mathlib/RingTheory/Norm/Basic.lean index 2d648f8fba650..b5bb4c471ca72 100644 --- a/Mathlib/RingTheory/Norm/Basic.lean +++ b/Mathlib/RingTheory/Norm/Basic.lean @@ -286,44 +286,6 @@ lemma norm_eq_of_equiv_equiv {A₁ B₁ A₂ B₂ : Type*} [CommRing A₁] [Ring simp only [RingHom.coe_comp, RingHom.coe_coe, Function.comp_apply, _root_.map_mul, RingEquiv.symm_apply_apply, commutes] -variable {F} (L) - --- TODO. Generalize this proof to rings -/--Let $F / L / K$ be a tower of finite extensions of fields. Then -$\text{Norm}_{F/K} = \text{Norm}_{L/K} \circ \text{Norm}_{F/L}$.-/ -@[stacks 0BIJ "This is a special case of 0BIJ norm, we require separability here. "] -theorem norm_norm [Algebra L F] [IsScalarTower K L F] [Algebra.IsSeparable K F] (x : F) : - norm K (norm L x) = norm K x := by - by_cases hKF : FiniteDimensional K F - · let A := AlgebraicClosure K - apply (algebraMap K A).injective - haveI : FiniteDimensional L F := FiniteDimensional.right K L F - haveI : FiniteDimensional K L := FiniteDimensional.left K L F - haveI : Algebra.IsSeparable K L := Algebra.isSeparable_tower_bot_of_isSeparable K L F - haveI : Algebra.IsSeparable L F := Algebra.isSeparable_tower_top_of_isSeparable K L F - letI : ∀ σ : L →ₐ[K] A, - haveI := σ.toRingHom.toAlgebra - Fintype (F →ₐ[L] A) := fun _ => inferInstance - rw [norm_eq_prod_embeddings K A (_ : F), - Fintype.prod_equiv algHomEquivSigma (fun σ : F →ₐ[K] A => σ x) - (fun π : Σ f : L →ₐ[K] A, _ => (π.2 : F → A) x) fun _ => rfl] - suffices ∀ σ : L →ₐ[K] A, - haveI := σ.toRingHom.toAlgebra - ∏ π : F →ₐ[L] A, π x = σ (norm L x) - by simp_rw [← Finset.univ_sigma_univ, Finset.prod_sigma, this, norm_eq_prod_embeddings] - intro σ - letI : Algebra L A := σ.toRingHom.toAlgebra - rw [← norm_eq_prod_embeddings L A (_ : F)] - simp [RingHom.algebraMap_toAlgebra] - · rw [norm_eq_one_of_not_module_finite hKF] - by_cases hKL : FiniteDimensional K L - · have hLF : ¬FiniteDimensional L F := by - refine (mt ?_) hKF - intro hKF - exact FiniteDimensional.trans K L F - rw [norm_eq_one_of_not_module_finite hLF, _root_.map_one] - · rw [norm_eq_one_of_not_module_finite hKL] - end EqProdEmbeddings end Algebra diff --git a/Mathlib/RingTheory/Norm/Transitivity.lean b/Mathlib/RingTheory/Norm/Transitivity.lean new file mode 100644 index 0000000000000..7476a9d8ce06c --- /dev/null +++ b/Mathlib/RingTheory/Norm/Transitivity.lean @@ -0,0 +1,196 @@ +/- +Copyright (c) 2024 Anne Baanen. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Junyan Xu +-/ +import Mathlib.LinearAlgebra.Matrix.Block +import Mathlib.LinearAlgebra.Matrix.Charpoly.Coeff +import Mathlib.RingTheory.Norm.Defs +import Mathlib.RingTheory.PolynomialAlgebra + +/-! +# Transitivity of algebra norm + +Suppose we have an `R`-algebra `S` with a finite basis. For each `s : S`, +the determinant of the linear map given by multiplying by `s` gives information +about the roots of the minimal polynomial of `s` over `R`. + +## References + + * [silvester2000] Silvester, *Determinants of Block Matrices*, The Mathematical Gazette (2000). + +-/ + +variable {R S A n m : Type*} [CommRing R] [CommRing S] +variable (M : Matrix m m S) [DecidableEq m] [DecidableEq n] (k : m) +open Matrix Polynomial + +namespace Algebra.Norm.Transitivity + +/-- Given a ((m-1)+1)x((m-1)+1) block matrix `M = [[A,b],[c,d]]`, `auxMat M k` is the auxiliary +matrix `[[dI,0],[-c,1]]`. `k` corresponds to the last row/column of the matrix. -/ +def auxMat : Matrix m m S := + of fun i j ↦ + if j = k then + if i = k then 1 else 0 + else if i = k then -M k j + else if i = j then M k k + else 0 + +/-- `aux M k` is lower triangular. -/ +lemma auxMat_blockTriangular : (auxMat M k).BlockTriangular (· ≠ k) := + fun i j lt ↦ by + simp_rw [lt_iff_not_le, le_Prop_eq, Classical.not_imp, not_not] at lt + rw [auxMat, of_apply, if_pos lt.2, if_neg lt.1] + +lemma auxMat_toSquareBlock_ne : (auxMat M k).toSquareBlock (· ≠ k) True = M k k • 1 := by + ext i j + simp [auxMat, toSquareBlock_def, if_neg (of_eq_true i.2), if_neg (of_eq_true j.2), + Matrix.one_apply, Subtype.ext_iff] + +lemma auxMat_toSquareBlock_eq : (auxMat M k).toSquareBlock (· ≠ k) False = 1 := by + ext ⟨i, hi⟩ ⟨j, hj⟩ + rw [eq_iff_iff, iff_false, not_not] at hi hj + simp [auxMat, toSquareBlock_def, if_pos hi, if_pos hj, Matrix.one_apply, if_pos (hj ▸ hi)] + +variable [Fintype m] + +/-- `M * aux M k` is upper triangular. -/ +lemma mul_auxMat_blockTriangular : (M * auxMat M k).BlockTriangular (· = k) := + fun i j lt ↦ by + simp_rw [lt_iff_not_le, le_Prop_eq, Classical.not_imp] at lt + simp_rw [Matrix.mul_apply, auxMat, of_apply, if_neg lt.2, mul_ite, mul_neg, mul_zero] + rw [Finset.sum_ite, Finset.filter_eq', if_pos (Finset.mem_univ _), Finset.sum_singleton, + Finset.sum_ite_eq', if_pos, lt.1, mul_comm, neg_add_cancel] + exact Finset.mem_filter.mpr ⟨Finset.mem_univ _, lt.2⟩ + +/-- The lower-right corner of `M * aux M k` is the same as the corner of `M`. -/ +lemma mul_auxMat_corner : (M * auxMat M k) k k = M k k := by simp [Matrix.mul_apply, auxMat] + +lemma mul_auxMat_toSquareBlock_eq : + (M * auxMat M k).toSquareBlock (· = k) True = M k k • 1 := by + ext ⟨i, hi⟩ ⟨j, hj⟩ + rw [eq_iff_iff, iff_true] at hi hj + simp [toSquareBlock_def, hi, hj, mul_auxMat_corner] + +set_option quotPrecheck false in +/-- The upper-left block of `M * aux M k`. -/ +notation "mulAuxMatBlock" => (M * auxMat M k).toSquareBlock (· = k) False + +lemma det_mul_corner_pow : + M.det * M k k ^ (Fintype.card m - 1) = M k k * (mulAuxMatBlock).det := by + trans (M * auxMat M k).det + · simp [det_mul, (auxMat_blockTriangular M k).det_fintype, + auxMat_toSquareBlock_ne, auxMat_toSquareBlock_eq] + rw [(mul_auxMat_blockTriangular M k).det_fintype, Fintype.prod_Prop, mul_auxMat_toSquareBlock_eq] + simp_rw [det_smul_of_tower, eq_iff_iff, iff_true, Fintype.card_unique, + pow_one, det_one, smul_eq_mul, mul_one] + -- `Decidable (P = Q)` diamond induced by `Prop.linearOrder`, which is classical, when `P` and `Q` + -- are themselves decidable. + convert rfl + +/-- A matrix with X added to the corner. -/ +noncomputable def cornerAddX : Matrix m m S[X] := + (diagonal fun i ↦ if i = k then X else 0) + C.mapMatrix M + +variable [Fintype n] (f : S →+* Matrix n n R) + +lemma polyToMatrix_cornerAddX : + f.polyToMatrix (cornerAddX M k k k) = (-f (M k k)).charmatrix := by + simp [cornerAddX, Matrix.add_apply, charmatrix, + RingHom.polyToMatrix, ← AlgEquiv.symm_toRingEquiv, map_neg] + +lemma eval_zero_det_det : eval 0 (f.polyToMatrix (cornerAddX M k).det).det = (f M.det).det := by + rw [← coe_evalRingHom, RingHom.map_det, ← RingHom.comp_apply, + evalRingHom_mapMatrix_comp_polyToMatrix, f.comp_apply, RingHom.map_det] + congr; ext; simp [cornerAddX, diagonal, apply_ite] + +lemma eval_zero_comp_det : + eval 0 (comp m m n n R[X] <| f.polyToMatrix.mapMatrix <| cornerAddX M k).det = + (comp m m n n R <| f.mapMatrix M).det := by + simp_rw [← coe_evalRingHom, RingHom.map_det, ← compRingEquiv_apply, ← RingEquiv.coe_toRingHom, + ← RingHom.comp_apply, ← RingHom.comp_assoc, evalRingHom_mapMatrix_comp_compRingEquiv, + RingHom.comp_assoc, RingHom.mapMatrix_comp, evalRingHom_mapMatrix_comp_polyToMatrix, + ← RingHom.mapMatrix_comp, RingHom.comp_apply] + congr with i j + simp [cornerAddX, diagonal, apply_ite] + +theorem comp_det_mul_pow : + ((f.mapMatrix M).comp m m n n R).det * (f (M k k)).det ^ (Fintype.card m - 1) = + (f (M k k)).det * ((f.mapMatrix mulAuxMatBlock).comp _ _ n n R).det := by + trans ((f.mapMatrix (M * auxMat M k)).compRingEquiv m n R).det + · simp_rw [_root_.map_mul, det_mul, f.mapMatrix_apply, compRingEquiv_apply, + ((auxMat_blockTriangular M k).map f).comp.det_fintype, Fintype.prod_Prop, + comp_toSquareBlock (b := (· ≠ k)), det_reindex_self, map_toSquareBlock, + auxMat_toSquareBlock_eq, auxMat_toSquareBlock_ne, smul_one_eq_diagonal, ← diagonal_one, + diagonal_map (map_zero _), comp_diagonal, det_reindex_self] + simp + · simp_rw [f.mapMatrix_apply, compRingEquiv_apply, ((mul_auxMat_blockTriangular M k).map f).comp + |>.det_fintype, Fintype.prod_Prop, comp_toSquareBlock (b := (· = k)), det_reindex_self, + map_toSquareBlock, mul_auxMat_toSquareBlock_eq, smul_one_eq_diagonal, + diagonal_map (map_zero _), comp_diagonal, det_reindex_self] + simp + +variable {M f} in +lemma det_det_aux + (ih : ∀ M, (f M.det).det = ((f.mapMatrix M).comp {a // (a = k) = False} _ n n R).det) : + ((f M.det).det - ((f.mapMatrix M).comp m m n n R).det) * + (f (M k k)).det ^ (Fintype.card m - 1) = 0 := by + rw [sub_mul, comp_det_mul_pow, ← det_pow, ← map_pow, ← det_mul, ← _root_.map_mul, + det_mul_corner_pow, _root_.map_mul, det_mul, ih, sub_self] + +end Algebra.Norm.Transitivity + +open Algebra.Norm.Transitivity + +/-- The main result in Silvester's paper *Determinants of Block Matrices*: the determinant of +a block matrix with commuting, equal-sized, square blocks can be computed by taking determinants +twice in a row: first take the determinant over the commutative ring generated by the +blocks (`S` here), then take the determinant over the base ring. -/ +theorem Matrix.det_det [Fintype m] [Fintype n] (f : S →+* Matrix n n R) : + (f M.det).det = ((f.mapMatrix M).comp m m n n R).det := by + set l := Fintype.card m with hl + clear_value l; revert R S m + induction' l with l ih <;> intro R S m _ _ M _ _ f card + · rw [eq_comm, Fintype.card_eq_zero_iff] at card + simp_rw [Matrix.det_isEmpty, _root_.map_one, det_one] + have ⟨k⟩ := Fintype.card_pos_iff.mp (l.succ_pos.trans_eq card) + let f' := f.polyToMatrix + let M' := cornerAddX M k + have : (f' M'.det).det = ((f'.mapMatrix M').comp m m n n R[X]).det := by + refine sub_eq_zero.mp <| mem_nonZeroDivisors_iff.mp + (pow_mem ?_ _) _ (det_det_aux k fun M ↦ ih _ _ <| by simp [← card]) + rw [polyToMatrix_cornerAddX, ← charpoly] + exact (Matrix.charpoly_monic _).mem_nonZeroDivisors + rw [← eval_zero_det_det, congr_arg (eval 0) this, eval_zero_comp_det] + +variable [Algebra R S] [Module.Free R S] + +theorem LinearMap.det_restrictScalars [AddCommGroup A] [Module R A] [Module S A] + [IsScalarTower R S A] [Module.Free S A] {f : A →ₗ[S] A} : + (f.restrictScalars R).det = Algebra.norm R f.det := by + nontriviality R + cases subsingleton_or_nontrivial A + · simp_rw [det_eq_one_of_subsingleton, _root_.map_one] + have := Module.nontrivial S A + let ⟨ιS, bS⟩ := Module.Free.exists_basis (R := R) (M := S) + let ⟨ιA, bA⟩ := Module.Free.exists_basis (R := S) (M := A) + have := bS.index_nonempty + have := bA.index_nonempty + cases fintypeOrInfinite ιS; swap + · rw [Algebra.norm_eq_one_of_not_module_finite (Module.not_finite_of_infinite_basis bS), + det_eq_one_of_not_module_finite (Module.not_finite_of_infinite_basis (bS.smulTower bA))] + cases fintypeOrInfinite ιA; swap + · rw [det_eq_one_of_not_module_finite (Module.not_finite_of_infinite_basis bA), _root_.map_one, + det_eq_one_of_not_module_finite (Module.not_finite_of_infinite_basis (bS.smulTower bA))] + classical + rw [Algebra.norm_eq_matrix_det bS, ← AlgHom.coe_toRingHom, ← det_toMatrix bA, det_det, + ← det_toMatrix (bS.smulTower' bA), restrictScalars_toMatrix] + rfl + +/--Let A/S/R be a tower of finite free tower of rings (with R and S commutative). +Then $\text{Norm}_{A/R} = \text{Norm}_{A/S} \circ \text{Norm}_{S/R}$.-/ +theorem Algebra.norm_norm {A} [Ring A] [Algebra R A] [Algebra S A] + [IsScalarTower R S A] [Module.Free S A] {a : A} : + norm R (norm S a) = norm R a := by + rw [norm_apply S, norm_apply R a, ← LinearMap.det_restrictScalars]; rfl diff --git a/Mathlib/RingTheory/PolynomialAlgebra.lean b/Mathlib/RingTheory/PolynomialAlgebra.lean index 278543baf82c4..545734d52795a 100644 --- a/Mathlib/RingTheory/PolynomialAlgebra.lean +++ b/Mathlib/RingTheory/PolynomialAlgebra.lean @@ -5,6 +5,7 @@ Authors: Kim Morrison -/ import Mathlib.Algebra.Polynomial.AlgebraMap import Mathlib.Data.Matrix.Basis +import Mathlib.Data.Matrix.Composition import Mathlib.Data.Matrix.DMatrix import Mathlib.RingTheory.MatrixAlgebra @@ -285,3 +286,19 @@ theorem support_subset_support_matPolyEquiv (m : Matrix n n R[X]) (i j : n) : intro hk rw [← matPolyEquiv_coeff_apply, hk] rfl + +variable {A} +/-- Extend a ring hom `A → Mₙ(R)` to a ring hom `A[X] → Mₙ(R[X])`. -/ +def RingHom.polyToMatrix (f : A →+* Matrix n n R) : A[X] →+* Matrix n n R[X] := + matPolyEquiv.symm.toRingHom.comp (mapRingHom f) + +variable {S : Type*} [CommSemiring S] (f : S →+* Matrix n n R) + +lemma evalRingHom_mapMatrix_comp_polyToMatrix : + (evalRingHom 0).mapMatrix.comp f.polyToMatrix = f.comp (evalRingHom 0) := by + ext <;> simp [RingHom.polyToMatrix, ← AlgEquiv.symm_toRingEquiv, diagonal, apply_ite] + +lemma evalRingHom_mapMatrix_comp_compRingEquiv {m} [Fintype m] [DecidableEq m] : + (evalRingHom 0).mapMatrix.comp (compRingEquiv m n R[X]) = + (compRingEquiv m n R).toRingHom.comp (evalRingHom 0).mapMatrix.mapMatrix := by + ext; simp diff --git a/docs/references.bib b/docs/references.bib index d397a999563cc..e42ea24949028 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -3304,6 +3304,19 @@ @Book{ silverman2009 year = {2009} } +@Article{ silvester2000, + issn = {00255572}, + url = {http://www.jstor.org/stable/3620776}, + author = {Silvester, John R.}, + journal = {The Mathematical Gazette}, + number = {501}, + pages = {460--467}, + publisher = {Mathematical Association}, + title = {Determinants of Block Matrices}, + volume = {84}, + year = {2000} +} + @Book{ simon2011, author = {Simon, Barry}, title = {Convexity: An Analytic Viewpoint}, From 625ac34a53dfa0eca4fb8a3a933409238e4f43f0 Mon Sep 17 00:00:00 2001 From: Daniel Weber Date: Mon, 16 Dec 2024 15:39:26 +0000 Subject: [PATCH 747/829] =?UTF-8?q?feat(RingTheory/Valuation/Basic):=20API?= =?UTF-8?q?=20for=20conversion=20`AddValuation`=20=E2=86=94=20`Valuation`?= =?UTF-8?q?=20(#18786)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- .../Order/GroupWithZero/Canonical.lean | 48 ++++++----- Mathlib/RingTheory/Valuation/Basic.lean | 80 +++++++++++++++++-- 2 files changed, 105 insertions(+), 23 deletions(-) diff --git a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean index d26601f412fd7..5c4e16d2e9c36 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean @@ -82,11 +82,16 @@ theorem zero_lt_iff : 0 < a ↔ a ≠ 0 := theorem ne_zero_of_lt (h : b < a) : a ≠ 0 := fun h1 ↦ not_lt_zero' <| show b < 0 from h1 ▸ h instance instLinearOrderedAddCommMonoidWithTopAdditiveOrderDual : - LinearOrderedAddCommMonoidWithTop (Additive αᵒᵈ) := - { Additive.orderedAddCommMonoid, Additive.linearOrder with - top := (0 : α) - top_add' := fun a ↦ zero_mul a.toMul - le_top := fun _ ↦ zero_le' } + LinearOrderedAddCommMonoidWithTop (Additive αᵒᵈ) where + top := (0 : α) + top_add' := fun a ↦ zero_mul a.toMul + le_top := fun _ ↦ zero_le' + +instance instLinearOrderedAddCommMonoidWithTopOrderDualAdditive : + LinearOrderedAddCommMonoidWithTop (Additive α)ᵒᵈ where + top := OrderDual.toDual (Additive.ofMul 0) + top_add' := fun a ↦ zero_mul (Additive.toMul (OrderDual.ofDual a)) + le_top := fun a ↦ @zero_le' _ _ (Additive.toMul (OrderDual.ofDual a)) variable [NoZeroDivisors α] @@ -191,11 +196,17 @@ theorem OrderIso.mulRight₀'_symm {a : α} (ha : a ≠ 0) : ext rfl -instance : LinearOrderedAddCommGroupWithTop (Additive αᵒᵈ) := - { Additive.subNegMonoid, instLinearOrderedAddCommMonoidWithTopAdditiveOrderDual, - Additive.instNontrivial with - neg_top := inv_zero (G₀ := α) - add_neg_cancel := fun a ha ↦ mul_inv_cancel₀ (G₀ := α) (id ha : a.toMul ≠ 0) } +instance : LinearOrderedAddCommGroupWithTop (Additive αᵒᵈ) where + __ := Additive.subNegMonoid + __ := instLinearOrderedAddCommMonoidWithTopAdditiveOrderDual + neg_top := inv_zero (G₀ := α) + add_neg_cancel := fun a ha ↦ mul_inv_cancel₀ (G₀ := α) (id ha : a.toMul ≠ 0) + +instance : LinearOrderedAddCommGroupWithTop (Additive α)ᵒᵈ where + __ := instSubNegAddMonoidOrderDual + __ := instLinearOrderedAddCommMonoidWithTopOrderDualAdditive + neg_top := inv_zero (G₀ := α) + add_neg_cancel := fun a ha ↦ mul_inv_cancel₀ (G₀ := α) (id ha : a.toMul ≠ 0) @[deprecated pow_lt_pow_right₀ (since := "2024-11-18")] lemma pow_lt_pow_succ (ha : 1 < a) : a ^ n < a ^ n.succ := pow_lt_pow_right₀ ha n.lt_succ_self @@ -204,14 +215,15 @@ end LinearOrderedCommGroupWithZero instance instLinearOrderedCommMonoidWithZeroMultiplicativeOrderDual [LinearOrderedAddCommMonoidWithTop α] : - LinearOrderedCommMonoidWithZero (Multiplicative αᵒᵈ) := - { Multiplicative.orderedCommMonoid, Multiplicative.linearOrder with - zero := Multiplicative.ofAdd (OrderDual.toDual ⊤) - zero_mul := @top_add _ (_) - -- Porting note: Here and elsewhere in the file, just `zero_mul` worked in Lean 3. See - -- https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Type.20synonyms - mul_zero := @add_top _ (_) - zero_le_one := (le_top : (0 : α) ≤ ⊤) } + LinearOrderedCommMonoidWithZero (Multiplicative αᵒᵈ) where + __ := Multiplicative.orderedCommMonoid + __ := Multiplicative.linearOrder + zero := Multiplicative.ofAdd (OrderDual.toDual ⊤) + zero_mul := @top_add _ (_) + -- Porting note: Here and elsewhere in the file, just `zero_mul` worked in Lean 3. See + -- https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Type.20synonyms + mul_zero := @add_top _ (_) + zero_le_one := (le_top : (0 : α) ≤ ⊤) @[simp] theorem ofAdd_toDual_eq_zero_iff [LinearOrderedAddCommMonoidWithTop α] diff --git a/Mathlib/RingTheory/Valuation/Basic.lean b/Mathlib/RingTheory/Valuation/Basic.lean index fe79b5797d6c9..ea3d9cfaa4ec3 100644 --- a/Mathlib/RingTheory/Valuation/Basic.lean +++ b/Mathlib/RingTheory/Valuation/Basic.lean @@ -635,11 +635,41 @@ theorem of_apply : (of f h0 h1 hadd hmul) r = f r := rfl /-- The `Valuation` associated to an `AddValuation` (useful if the latter is constructed using `AddValuation.of`). -/ -def valuation : Valuation R (Multiplicative Γ₀ᵒᵈ) := - v +def toValuation : AddValuation R Γ₀ ≃ Valuation R (Multiplicative Γ₀ᵒᵈ) := + Equiv.refl _ + +@[deprecated (since := "2024-11-09")] +alias valuation := toValuation + +/-- The `AddValuation` associated to a `Valuation`. +-/ +def ofValuation : Valuation R (Multiplicative Γ₀ᵒᵈ) ≃ AddValuation R Γ₀ := + Equiv.refl _ + +@[simp] +lemma ofValuation_symm_eq : ofValuation.symm = toValuation (R := R) (Γ₀ := Γ₀) := rfl @[simp] -theorem valuation_apply (r : R) : v.valuation r = Multiplicative.ofAdd (OrderDual.toDual (v r)) := +lemma toValuation_symm_eq : toValuation.symm = ofValuation (R := R) (Γ₀ := Γ₀) := rfl + +@[simp] +lemma ofValuation_toValuation : ofValuation (toValuation v) = v := rfl + +@[simp] +lemma toValuation_ofValuation (v : Valuation R (Multiplicative Γ₀ᵒᵈ)) : + toValuation (ofValuation v) = v := rfl + +@[simp] +theorem toValuation_apply (r : R) : + toValuation v r = Multiplicative.ofAdd (OrderDual.toDual (v r)) := + rfl + +@[deprecated (since := "2024-11-09")] +alias valuation_apply := toValuation_apply + +@[simp] +theorem ofValuation_apply (v : Valuation R (Multiplicative Γ₀ᵒᵈ)) (r : R) : + ofValuation v r = OrderDual.ofDual (Multiplicative.toAdd (v r)) := rfl end @@ -751,11 +781,11 @@ variable [LinearOrderedAddCommGroupWithTop Γ₀] [Ring R] (v : AddValuation R @[simp] theorem map_inv (v : AddValuation K Γ₀) {x : K} : v x⁻¹ = - (v x) := - map_inv₀ v.valuation x + map_inv₀ (toValuation v) x @[simp] theorem map_div (v : AddValuation K Γ₀) {x y : K} : v (x / y) = v x - v y := - map_div₀ v.valuation x y + map_div₀ (toValuation v) x y @[simp] theorem map_neg (x : R) : v (-x) = v x := @@ -859,3 +889,43 @@ end Supp -- end of section end AddValuation + +namespace Valuation + + +variable {Γ₀ : Type*} [Ring R] [LinearOrderedCommMonoidWithZero Γ₀] + +/-- The `AddValuation` associated to a `Valuation`. -/ +def toAddValuation : Valuation R Γ₀ ≃ AddValuation R (Additive Γ₀)ᵒᵈ := + AddValuation.ofValuation (R := R) (Γ₀ := (Additive Γ₀)ᵒᵈ) + +/-- The `Valuation` associated to a `AddValuation`. +-/ +def ofAddValuation : AddValuation R (Additive Γ₀)ᵒᵈ ≃ Valuation R Γ₀ := + AddValuation.toValuation + +@[simp] +lemma ofAddValuation_symm_eq : ofAddValuation.symm = toAddValuation (R := R) (Γ₀ := Γ₀) := rfl + +@[simp] +lemma toAddValuation_symm_eq : toAddValuation.symm = ofAddValuation (R := R) (Γ₀ := Γ₀) := rfl + +@[simp] +lemma ofAddValuation_toAddValuation (v : Valuation R Γ₀) : + ofAddValuation (toAddValuation v) = v := rfl + +@[simp] +lemma toValuation_ofValuation (v : AddValuation R (Additive Γ₀)ᵒᵈ) : + toAddValuation (ofAddValuation v) = v := rfl + +@[simp] +theorem toAddValuation_apply (v : Valuation R Γ₀) (r : R) : + toAddValuation v r = OrderDual.toDual (Additive.ofMul (v r)) := + rfl + +@[simp] +theorem ofAddValuation_apply (v : AddValuation R (Additive Γ₀)ᵒᵈ) (r : R) : + ofAddValuation v r = Additive.toMul (OrderDual.ofDual (v r)) := + rfl + +end Valuation From 3a32603f07dfe9b7acd6aa0f967844204e2e341d Mon Sep 17 00:00:00 2001 From: Paul Reichert <6992158+datokrat@users.noreply.github.com> Date: Mon, 16 Dec 2024 16:45:52 +0000 Subject: [PATCH 748/829] feat(CategoryTheory/Abelian/GrothendieckAxioms): introduce Grothendieck categories (#18510) Define Grothendieck categories in a new separate file. According to our definition, a Grothendieck category w.r.t an universe w is an abelian category that is locally small w.r.t w, has exact filtered colimits of size `w` (AB5) and has a separator. Some essential theorems are provided: * The definition is invariant under equivalences of categories, and in particular, under shrinking the hom sets to size `w`. * Grothendieck categories are have all limits and colimits of size w (completeness and cocompleteness). In order to facilitate the proofs of the theorems, the new lemma `hasLimits_of_hasColimits_of_hasSeparator` (a variant of `hasLimits_of_hasColimits_of_isSeparating`) and its dual is introduced in `Adjunction.AdjointFunctorTheorems`. The module `Abelian.Transfer` is extended with the induced preadditive and abelian structure of `ShrinkHoms` and existing transfer lemmas are generalized so that the hom sets need not have the same universes. Transfer lemmas for separators and coseparators are added to `Adjunction.Generator`. Moreover, some notable changes are made in `Abelian.GrothendieckAxioms.Basic`. Besides adding `@[stacks]` attributes, `hasExactColimitsOfShape_of_equiv` is renamed to `HasExactColimitsOfShape.of_domain_equivalence` and is afforded a counterpart `HasExactColimitsOfShape.of_codomain_equivalence` (with similar changes to `HasExactLimitsOfShape`). Co-authored-by: Paul Reichert --- Mathlib.lean | 1 + .../Abelian/GrothendieckAxioms/Basic.lean | 53 ++++++--- .../Abelian/GrothendieckCategory.lean | 105 ++++++++++++++++++ Mathlib/CategoryTheory/Abelian/Transfer.lean | 77 +++++++++++-- .../Adjunction/AdjointFunctorTheorems.lean | 18 ++- Mathlib/CategoryTheory/EssentiallySmall.lean | 3 + Mathlib/CategoryTheory/Generator.lean | 42 +++++++ 7 files changed, 274 insertions(+), 25 deletions(-) create mode 100644 Mathlib/CategoryTheory/Abelian/GrothendieckCategory.lean diff --git a/Mathlib.lean b/Mathlib.lean index 00329c87c9e80..d3490bf12958a 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1575,6 +1575,7 @@ import Mathlib.CategoryTheory.Abelian.Generator import Mathlib.CategoryTheory.Abelian.GrothendieckAxioms.Basic import Mathlib.CategoryTheory.Abelian.GrothendieckAxioms.FunctorCategory import Mathlib.CategoryTheory.Abelian.GrothendieckAxioms.Sheaf +import Mathlib.CategoryTheory.Abelian.GrothendieckCategory import Mathlib.CategoryTheory.Abelian.Images import Mathlib.CategoryTheory.Abelian.Injective import Mathlib.CategoryTheory.Abelian.InjectiveResolution diff --git a/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms/Basic.lean b/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms/Basic.lean index 251628d72e070..b96f54ec72e3b 100644 --- a/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms/Basic.lean +++ b/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms/Basic.lean @@ -18,7 +18,7 @@ basic facts about them. ## Definitions -- `HasExactColimitsOfShape J` -- colimits of shape `J` are exact. +- `HasExactColimitsOfShape J C` -- colimits of shape `J` in `C` are exact. - The dual of the above definitions, called `HasExactLimitsOfShape`. - `AB4` -- coproducts are exact (this is formulated in terms of `HasExactColimitsOfShape`). - `AB5` -- filtered colimits are exact (this is formulated in terms of `HasExactColimitsOfShape`). @@ -26,6 +26,9 @@ basic facts about them. ## Theorems - The implication from `AB5` to `AB4` is established in `AB4.ofAB5`. +- That `HasExactColimitsOfShape J C` is invariant under equivalences in both parameters is shown +in `HasExactColimitsOfShape.of_domain_equivalence` and +`HasExactColimitsOfShape.of_codomain_equivalence`. ## Remarks @@ -39,10 +42,6 @@ We do not require `Abelian` in the definition of `AB4` and `AB5` because these c individual axioms. An `AB4` category is an _abelian_ category satisfying `AB4`, and similarly for `AB5`. -## Projects - -- Add additional axioms, especially define Grothendieck categories. - ## References * [Stacks: Grothendieck's AB conditions](https://stacks.math.columbia.edu/tag/079A) @@ -127,26 +126,48 @@ Transport a `HasExactColimitsOfShape` along an equivalence of the shape. Note: When `C` has finite limits, this lemma holds with the equivalence replaced by a final functor, see `hasExactColimitsOfShape_of_final` below. -/ -lemma hasExactColimitsOfShape_of_equiv {J J' : Type*} [Category J] [Category J'] (e : J ≌ J') - [HasColimitsOfShape J C] [HasExactColimitsOfShape J C] : +lemma HasExactColimitsOfShape.of_domain_equivalence {J J' : Type*} [Category J] [Category J'] + (e : J ≌ J') [HasColimitsOfShape J C] [HasExactColimitsOfShape J C] : haveI : HasColimitsOfShape J' C := hasColimitsOfShape_of_equivalence e HasExactColimitsOfShape J' C := haveI : HasColimitsOfShape J' C := hasColimitsOfShape_of_equivalence e ⟨preservesFiniteLimits_of_natIso (Functor.Final.colimIso e.functor)⟩ +variable {C} in +lemma HasExactColimitsOfShape.of_codomain_equivalence (J : Type*) [Category J] {D : Type*} + [Category D] (e : C ≌ D) [HasColimitsOfShape J C] [HasExactColimitsOfShape J C] : + haveI : HasColimitsOfShape J D := Adjunction.hasColimitsOfShape_of_equivalence e.inverse + HasExactColimitsOfShape J D := by + haveI : HasColimitsOfShape J D := Adjunction.hasColimitsOfShape_of_equivalence e.inverse + refine ⟨⟨fun _ _ _ => ⟨@fun K => ?_⟩⟩⟩ + refine preservesLimit_of_natIso K (?_ : e.congrRight.inverse ⋙ colim ⋙ e.functor ≅ colim) + apply e.symm.congrRight.fullyFaithfulFunctor.preimageIso + exact isoWhiskerLeft (_ ⋙ colim) e.unitIso.symm ≪≫ (preservesColimitNatIso e.inverse).symm + /-- Transport a `HasExactLimitsOfShape` along an equivalence of the shape. Note: When `C` has finite colimits, this lemma holds with the equivalence replaced by a initial functor, see `hasExactLimitsOfShape_of_initial` below. -/ -lemma hasExactLimitsOfShape_of_equiv {J J' : Type*} [Category J] [Category J'] (e : J ≌ J') - [HasLimitsOfShape J C] [HasExactLimitsOfShape J C] : +lemma HasExactLimitsOfShape.of_domain_equivalence {J J' : Type*} [Category J] [Category J'] + (e : J ≌ J') [HasLimitsOfShape J C] [HasExactLimitsOfShape J C] : haveI : HasLimitsOfShape J' C := hasLimitsOfShape_of_equivalence e HasExactLimitsOfShape J' C := haveI : HasLimitsOfShape J' C := hasLimitsOfShape_of_equivalence e ⟨preservesFiniteColimits_of_natIso (Functor.Initial.limIso e.functor)⟩ +variable {C} in +lemma HasExactLimitsOfShape.of_codomain_equivalence (J : Type*) [Category J] {D : Type*} + [Category D] (e : C ≌ D) [HasLimitsOfShape J C] [HasExactLimitsOfShape J C] : + haveI : HasLimitsOfShape J D := Adjunction.hasLimitsOfShape_of_equivalence e.inverse + HasExactLimitsOfShape J D := by + haveI : HasLimitsOfShape J D := Adjunction.hasLimitsOfShape_of_equivalence e.inverse + refine ⟨⟨fun _ _ _ => ⟨@fun K => ?_⟩⟩⟩ + refine preservesColimit_of_natIso K (?_ : e.congrRight.inverse ⋙ lim ⋙ e.functor ≅ lim) + apply e.symm.congrRight.fullyFaithfulFunctor.preimageIso + exact isoWhiskerLeft (_ ⋙ lim) e.unitIso.symm ≪≫ (preservesLimitNatIso e.inverse).symm + /-- A category `C` which has coproducts is said to have `AB4` of size `w` provided that coproducts of size `w` are exact. @@ -161,13 +182,14 @@ attribute [instance] AB4OfSize.ofShape A category `C` which has coproducts is said to have `AB4` provided that coproducts are exact. -/ +@[stacks 079B] abbrev AB4 [HasCoproducts C] := AB4OfSize.{v} C lemma AB4OfSize_shrink [HasCoproducts.{max w w'} C] [AB4OfSize.{max w w'} C] : haveI : HasCoproducts.{w} C := hasCoproducts_shrink.{w, w'} AB4OfSize.{w} C := haveI := hasCoproducts_shrink.{w, w'} (C := C) - ⟨fun J ↦ hasExactColimitsOfShape_of_equiv C + ⟨fun J ↦ HasExactColimitsOfShape.of_domain_equivalence C (Discrete.equivalence Equiv.ulift : Discrete (ULift.{w'} J) ≌ _)⟩ instance (priority := 100) [HasCoproducts.{w} C] [AB4OfSize.{w} C] : @@ -176,7 +198,7 @@ instance (priority := 100) [HasCoproducts.{w} C] [AB4OfSize.{w} C] : /-- A category `C` which has products is said to have `AB4Star` (in literature `AB4*`) provided that products are exact. -/ -@[pp_with_univ] +@[pp_with_univ, stacks 079B] class AB4StarOfSize [HasProducts.{w} C] where ofShape (α : Type w) : HasExactLimitsOfShape (Discrete α) C @@ -190,7 +212,7 @@ lemma AB4StarOfSize_shrink [HasProducts.{max w w'} C] [AB4StarOfSize.{max w w'} haveI : HasProducts.{w} C := hasProducts_shrink.{w, w'} AB4StarOfSize.{w} C := haveI := hasProducts_shrink.{w, w'} (C := C) - ⟨fun J ↦ hasExactLimitsOfShape_of_equiv C + ⟨fun J ↦ HasExactLimitsOfShape.of_domain_equivalence C (Discrete.equivalence Equiv.ulift : Discrete (ULift.{w'} J) ≌ _)⟩ instance (priority := 100) [HasProducts.{w} C] [AB4StarOfSize.{w} C] : @@ -236,6 +258,7 @@ attribute [instance] AB5OfSize.ofShape A category `C` which has filtered colimits is said to have `AB5` provided that filtered colimits are exact. -/ +@[stacks 079B] abbrev AB5 [HasFilteredColimits C] := AB5OfSize.{v, v} C lemma AB5OfSize_of_univLE [HasFilteredColimitsOfSize.{w₂, w₂'} C] [UnivLE.{w, w₂}] @@ -247,7 +270,7 @@ lemma AB5OfSize_of_univLE [HasFilteredColimitsOfSize.{w₂, w₂'} C] [UnivLE.{w intro J _ _ haveI := IsFiltered.of_equivalence ((ShrinkHoms.equivalence.{w₂} J).trans <| Shrink.equivalence.{w₂'} (ShrinkHoms.{w'} J)) - exact hasExactColimitsOfShape_of_equiv _ ((ShrinkHoms.equivalence.{w₂} J).trans <| + exact HasExactColimitsOfShape.of_domain_equivalence _ ((ShrinkHoms.equivalence.{w₂} J).trans <| Shrink.equivalence.{w₂'} (ShrinkHoms.{w'} J)).symm lemma AB5OfSize_shrink [HasFilteredColimitsOfSize.{max w w₂, max w' w₂'} C] @@ -260,7 +283,7 @@ lemma AB5OfSize_shrink [HasFilteredColimitsOfSize.{max w w₂, max w' w₂'} C] A category `C` which has cofiltered limits is said to have `AB5Star` (in literature `AB5*`) provided that cofiltered limits are exact. -/ -@[pp_with_univ] +@[pp_with_univ, stacks 079B] class AB5StarOfSize [HasCofilteredLimitsOfSize.{w, w'} C] where ofShape (J : Type w') [Category.{w} J] [IsCofiltered J] : HasExactLimitsOfShape J C @@ -281,7 +304,7 @@ lemma AB5StarOfSize_of_univLE [HasCofilteredLimitsOfSize.{w₂, w₂'} C] [UnivL intro J _ _ haveI := IsCofiltered.of_equivalence ((ShrinkHoms.equivalence.{w₂} J).trans <| Shrink.equivalence.{w₂'} (ShrinkHoms.{w'} J)) - exact hasExactLimitsOfShape_of_equiv _ ((ShrinkHoms.equivalence.{w₂} J).trans <| + exact HasExactLimitsOfShape.of_domain_equivalence _ ((ShrinkHoms.equivalence.{w₂} J).trans <| Shrink.equivalence.{w₂'} (ShrinkHoms.{w'} J)).symm lemma AB5StarOfSize_shrink [HasCofilteredLimitsOfSize.{max w w₂, max w' w₂'} C] diff --git a/Mathlib/CategoryTheory/Abelian/GrothendieckCategory.lean b/Mathlib/CategoryTheory/Abelian/GrothendieckCategory.lean new file mode 100644 index 0000000000000..3fbacaecba6e9 --- /dev/null +++ b/Mathlib/CategoryTheory/Abelian/GrothendieckCategory.lean @@ -0,0 +1,105 @@ +/- +Copyright (c) 2024 Paul Reichert. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Paul Reichert +-/ + +import Mathlib.CategoryTheory.Abelian.GrothendieckAxioms.Basic +import Mathlib.CategoryTheory.Abelian.Subobject +import Mathlib.CategoryTheory.Abelian.Transfer +import Mathlib.CategoryTheory.Adjunction.AdjointFunctorTheorems +import Mathlib.CategoryTheory.Limits.HasLimits + +/-! + +# Grothendieck categories + +This file defines Grothendieck categories and proves basic facts about them. + +## Definitions + +A Grothendieck category according to the Stacks project is an abelian category provided that it +has `AB5` and a separator. However, this definition is not invariant under equivalences of +categories. Therefore, if `C` is an abelian category, the class `IsGrothendieckAbelian.{w} C` has a +weaker definition that is also satisfied for categories that are merely equivalent to a +Grothendieck category in the former strict sense. + +## Theorems + +The invariance under equivalences of categories is established in +`IsGrothendieckAbelian.of_equivalence`. + +In particular, `ShrinkHoms.isGrothendieckAbelian C` shows that `ShrinkHoms C` satisfies our +definition of a Grothendieck category after shrinking its hom sets, which coincides with the strict +definition in this case. + +Relevant implications of `IsGrothendieckAbelian` are established in +`IsGrothendieckAbelian.hasLimits` and `IsGrothendieckAbelian.hasColimits`. + +## References + +* [Stacks: Grothendieck's AB conditions](https://stacks.math.columbia.edu/tag/079A) + +-/ + +namespace CategoryTheory + +open Limits + +universe w v u w₂ v₂ u₂ +variable (C : Type u) [Category.{v} C] (D : Type u₂) [Category.{v₂} D] + +/-- +If `C` is an abelian category, we shall say that it satisfies `IsGrothendieckAbelian.{w} C` +if it is locally small (relative to `w`), has exact filtered colimits of size `w` (AB5) and has a +separator. +If `[Category.{v} C]` and `w = v`, this means that `C` satisfies `AB5` and has a separator; +general results about Grothendieck abelian categories can be +reduced to this case using the instance `ShrinkHoms.isGrothendieckAbelian` below. + +The introduction of the auxiliary universe `w` shall be needed for certain +applications to categories of sheaves. That the present definition still preserves essential +properties of Grothendieck categories is ensured by `IsGrothendieckAbelian.of_equivalence`, +which shows that every instance for `C` implies an instance for `ShrinkHoms C` with hom sets in +`Type w`. +-/ +@[stacks 079B] +class IsGrothendieckAbelian [Abelian C] : Prop where + locallySmall : LocallySmall.{w} C := by infer_instance + hasFilteredColimitsOfSize : HasFilteredColimitsOfSize.{w, w} C := by infer_instance + ab5OfSize : AB5OfSize.{w, w} C := by infer_instance + hasSeparator : HasSeparator C := by infer_instance + +attribute [instance] IsGrothendieckAbelian.locallySmall + IsGrothendieckAbelian.hasFilteredColimitsOfSize IsGrothendieckAbelian.ab5OfSize + IsGrothendieckAbelian.hasSeparator + +variable {C} {D} in +theorem IsGrothendieckAbelian.of_equivalence [Abelian C] [Abelian D] + [IsGrothendieckAbelian.{w} C] (α : C ≌ D) : IsGrothendieckAbelian.{w} D := by + have hasFilteredColimits : HasFilteredColimitsOfSize.{w, w, v₂, u₂} D := + ⟨fun _ _ _ => Adjunction.hasColimitsOfShape_of_equivalence α.inverse⟩ + refine ⟨?_, hasFilteredColimits, ?_, ?_⟩ + · exact locallySmall_of_faithful α.inverse + · refine ⟨fun _ _ _ => ?_⟩ + exact HasExactColimitsOfShape.of_codomain_equivalence _ α + · exact HasSeparator.of_equivalence α + +instance ShrinkHoms.isGrothendieckAbelian [Abelian C] [IsGrothendieckAbelian.{w} C] : + IsGrothendieckAbelian.{w, w} (ShrinkHoms C) := + IsGrothendieckAbelian.of_equivalence <| ShrinkHoms.equivalence C + +section Instances + +variable [Abelian C] [IsGrothendieckAbelian.{w} C] + +instance IsGrothendieckAbelian.hasColimits : HasColimitsOfSize.{w, w} C := + has_colimits_of_finite_and_filtered + +instance IsGrothendieckAbelian.hasLimits : HasLimitsOfSize.{w, w} C := + have : HasLimits.{w, u} (ShrinkHoms C) := hasLimits_of_hasColimits_of_hasSeparator + Adjunction.has_limits_of_equivalence (ShrinkHoms.equivalence C |>.functor) + +end Instances + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Abelian/Transfer.lean b/Mathlib/CategoryTheory/Abelian/Transfer.lean index edda974b2f59c..0c4a50ffe85f5 100644 --- a/Mathlib/CategoryTheory/Abelian/Transfer.lean +++ b/Mathlib/CategoryTheory/Abelian/Transfer.lean @@ -4,8 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ import Mathlib.CategoryTheory.Abelian.Basic -import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Kernels import Mathlib.CategoryTheory.Adjunction.Limits +import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Kernels +import Mathlib.CategoryTheory.Preadditive.AdditiveFunctor +import Mathlib.Logic.Equiv.TransferInstance /-! # Transferring "abelian-ness" across a functor @@ -16,6 +18,9 @@ we have `F : C ⥤ D` `G : D ⥤ C` (both preserving zero morphisms), and further we have `adj : G ⊣ F` and `i : F ⋙ G ≅ 𝟭 C`, then `C` is also abelian. +A particular example is the transfer of `Abelian` instances from a category `C` to `ShrinkHoms C`; +see `ShrinkHoms.abelian`. In this case, we also transfer the `Preadditive` structure. + See ## Notes @@ -37,12 +42,12 @@ namespace CategoryTheory open Limits -universe v u₁ u₂ +universe v₁ v₂ u₁ u₂ namespace AbelianOfAdjunction -variable {C : Type u₁} [Category.{v} C] [Preadditive C] -variable {D : Type u₂} [Category.{v} D] [Abelian D] +variable {C : Type u₁} [Category.{v₁} C] [Preadditive C] +variable {D : Type u₂} [Category.{v₂} D] [Abelian D] variable (F : C ⥤ D) variable (G : D ⥤ C) [Functor.PreservesZeroMorphisms G] @@ -153,8 +158,8 @@ then `C` is also abelian. See -/ -def abelianOfAdjunction {C : Type u₁} [Category.{v} C] [Preadditive C] [HasFiniteProducts C] - {D : Type u₂} [Category.{v} D] [Abelian D] (F : C ⥤ D) [Functor.PreservesZeroMorphisms F] +def abelianOfAdjunction {C : Type u₁} [Category.{v₁} C] [Preadditive C] [HasFiniteProducts C] + {D : Type u₂} [Category.{v₂} D] [Abelian D] (F : C ⥤ D) [Functor.PreservesZeroMorphisms F] (G : D ⥤ C) [Functor.PreservesZeroMorphisms G] [PreservesFiniteLimits G] (i : F ⋙ G ≅ 𝟭 C) (adj : G ⊣ F) : Abelian C := by haveI := hasKernels F G i @@ -169,9 +174,65 @@ def abelianOfAdjunction {C : Type u₁} [Category.{v} C] [Preadditive C] [HasFin via a functor that preserves zero morphisms, then `C` is also abelian. -/ -def abelianOfEquivalence {C : Type u₁} [Category.{v} C] [Preadditive C] [HasFiniteProducts C] - {D : Type u₂} [Category.{v} D] [Abelian D] (F : C ⥤ D) [Functor.PreservesZeroMorphisms F] +def abelianOfEquivalence {C : Type u₁} [Category.{v₁} C] [Preadditive C] [HasFiniteProducts C] + {D : Type u₂} [Category.{v₂} D] [Abelian D] (F : C ⥤ D) [Functor.PreservesZeroMorphisms F] [F.IsEquivalence] : Abelian C := abelianOfAdjunction F F.inv F.asEquivalence.unitIso.symm F.asEquivalence.symm.toAdjunction +namespace ShrinkHoms + +universe w + +variable {C : Type*} [Category C] [LocallySmall.{w} C] + +section Preadditive + +variable [Preadditive C] + +noncomputable instance homGroup (P Q : ShrinkHoms C) : AddCommGroup (P ⟶ Q : Type w) := + Equiv.addCommGroup (equivShrink _).symm + +lemma functor_map_add {P Q : C} (f g : P ⟶ Q) : + (functor C).map (f + g) = + (functor C).map f + (functor C).map g := by + exact map_add (equivShrink.{w} (P ⟶ Q)).symm.addEquiv.symm f g + +lemma inverse_map_add {P Q : ShrinkHoms C} (f g : P ⟶ Q) : + (inverse C).map (f + g) = + (inverse C).map f + (ShrinkHoms.inverse C).map g := + map_add (equivShrink.{w} (P.fromShrinkHoms ⟶ Q.fromShrinkHoms)).symm.addEquiv f g + +variable (C) + +noncomputable instance preadditive : + Preadditive.{w} (ShrinkHoms C) where + homGroup := homGroup + add_comp _ _ _ _ _ _ := by + apply (inverse C).map_injective + simp only [inverse_map_add, Functor.map_comp, Preadditive.add_comp] + comp_add _ _ _ _ _ _ := by + apply (inverse C).map_injective + simp only [inverse_map_add, Functor.map_comp, Preadditive.comp_add] + +instance : (inverse C).Additive where + map_add := by apply inverse_map_add + +instance : (functor C).Additive where + map_add := by apply functor_map_add + +instance hasLimitsOfShape (J : Type*) [Category J] + [HasLimitsOfShape J C] : HasLimitsOfShape.{_, _, w} J (ShrinkHoms C) := + Adjunction.hasLimitsOfShape_of_equivalence (inverse C) + +instance hasFiniteLimits [HasFiniteLimits C] : + HasFiniteLimits.{w} (ShrinkHoms C) := ⟨fun _ => inferInstance⟩ + +end Preadditive + +variable (C) in +noncomputable instance abelian [Abelian C] : + Abelian.{w} (ShrinkHoms C) := abelianOfEquivalence (inverse C) + +end ShrinkHoms + end CategoryTheory diff --git a/Mathlib/CategoryTheory/Adjunction/AdjointFunctorTheorems.lean b/Mathlib/CategoryTheory/Adjunction/AdjointFunctorTheorems.lean index 6d42354bbaaa3..b4b03bb6b7eb7 100644 --- a/Mathlib/CategoryTheory/Adjunction/AdjointFunctorTheorems.lean +++ b/Mathlib/CategoryTheory/Adjunction/AdjointFunctorTheorems.lean @@ -29,9 +29,11 @@ This file also proves the special adjoint functor theorem, in the form: * If `G : D ⥤ C` preserves limits and `D` is complete, well-powered and has a small coseparating set, then `G` has a left adjoint: `isRightAdjointOfPreservesLimitsOfIsCoseparating` -Finally, we prove the following corollary of the special adjoint functor theorem: +Finally, we prove the following corollaries of the special adjoint functor theorem: * If `C` is complete, well-powered and has a small coseparating set, then it is cocomplete: - `hasColimits_of_hasLimits_of_isCoseparating` + `hasColimits_of_hasLimits_of_isCoseparating`, `hasColimits_of_hasLimits_of_hasCoseparator` +* If `C` is cocomplete, co-well-powered and has a small separating set, then it is complete: + `hasLimits_of_hasColimits_of_isSeparating`, `hasLimits_of_hasColimits_of_hasSeparator` -/ @@ -135,6 +137,18 @@ theorem hasLimits_of_hasColimits_of_isSeparating [HasColimits C] [WellPowered C hasLimitsOfShape_iff_isLeftAdjoint_const.2 (isLeftAdjoint_of_preservesColimits_of_isSeparating h𝒢 _) } +/-- A consequence of the special adjoint functor theorem: if `C` is complete, well-powered and + has a separator, then it is complete. -/ +theorem hasLimits_of_hasColimits_of_hasSeparator [HasColimits C] [HasSeparator C] + [WellPowered Cᵒᵖ] : HasLimits C := + hasLimits_of_hasColimits_of_isSeparating <| isSeparator_separator C + +/-- A consequence of the special adjoint functor theorem: if `C` is complete, well-powered and + has a coseparator, then it is cocomplete. -/ +theorem hasColimits_of_hasLimits_of_hasCoseparator [HasLimits C] [HasCoseparator C] + [WellPowered C] : HasColimits C := + hasColimits_of_hasLimits_of_isCoseparating <| isCoseparator_coseparator C + end Limits end CategoryTheory diff --git a/Mathlib/CategoryTheory/EssentiallySmall.lean b/Mathlib/CategoryTheory/EssentiallySmall.lean index 70c065ab24e7b..f041cca1723b9 100644 --- a/Mathlib/CategoryTheory/EssentiallySmall.lean +++ b/Mathlib/CategoryTheory/EssentiallySmall.lean @@ -174,6 +174,9 @@ noncomputable def equivalence : C ≌ ShrinkHoms C where unitIso := NatIso.ofComponents (fun _ ↦ Iso.refl _) counitIso := NatIso.ofComponents (fun _ ↦ Iso.refl _) +instance : (functor C).IsEquivalence := (equivalence C).isEquivalence_functor +instance : (inverse C).IsEquivalence := (equivalence C).isEquivalence_inverse + end ShrinkHoms namespace Shrink diff --git a/Mathlib/CategoryTheory/Generator.lean b/Mathlib/CategoryTheory/Generator.lean index 0e0311c2253f2..78a17790ce9fb 100644 --- a/Mathlib/CategoryTheory/Generator.lean +++ b/Mathlib/CategoryTheory/Generator.lean @@ -26,6 +26,7 @@ There are, of course, also the dual notions of coseparating and codetecting sets We * define predicates `IsSeparating`, `IsCoseparating`, `IsDetecting` and `IsCodetecting` on sets of objects; +* show that equivalences of categories preserves these notions; * show that separating and coseparating are dual notions; * show that detecting and codetecting are dual notions; * show that if `C` has equalizers, then detecting implies separating; @@ -80,6 +81,26 @@ def IsDetecting (𝒢 : Set C) : Prop := def IsCodetecting (𝒢 : Set C) : Prop := ∀ ⦃X Y : C⦄ (f : X ⟶ Y), (∀ G ∈ 𝒢, ∀ (h : X ⟶ G), ∃! h' : Y ⟶ G, f ≫ h' = h) → IsIso f +section Equivalence + +lemma IsSeparating.of_equivalence + {𝒢 : Set C} (h : IsSeparating 𝒢) {D : Type*} [Category D] (α : C ≌ D) : + IsSeparating (α.functor.obj '' 𝒢) := fun X Y f g H => + α.inverse.map_injective (h _ _ (fun Z hZ h => by + obtain ⟨h', rfl⟩ := (α.toAdjunction.homEquiv _ _).surjective h + simp only [Adjunction.homEquiv_unit, Category.assoc, ← Functor.map_comp, + H (α.functor.obj Z) (Set.mem_image_of_mem _ hZ) h'])) + +lemma IsCoseparating.of_equivalence + {𝒢 : Set C} (h : IsCoseparating 𝒢) {D : Type*} [Category D] (α : C ≌ D) : + IsCoseparating (α.functor.obj '' 𝒢) := fun X Y f g H => + α.inverse.map_injective (h _ _ (fun Z hZ h => by + obtain ⟨h', rfl⟩ := (α.symm.toAdjunction.homEquiv _ _).symm.surjective h + simp only [Adjunction.homEquiv_symm_apply, ← Category.assoc, ← Functor.map_comp, + Equivalence.symm_functor, H (α.functor.obj Z) (Set.mem_image_of_mem _ hZ) h'])) + +end Equivalence + section Dual theorem isSeparating_op_iff (𝒢 : Set C) : IsSeparating 𝒢.op ↔ IsCoseparating 𝒢 := by @@ -351,6 +372,17 @@ def IsDetector (G : C) : Prop := def IsCodetector (G : C) : Prop := IsCodetecting ({G} : Set C) + +section Equivalence + +theorem IsSeparator.of_equivalence {G : C} (h : IsSeparator G) (α : C ≌ D) : + IsSeparator (α.functor.obj G) := by simpa using IsSeparating.of_equivalence h α + +theorem IsCoseparator.of_equivalence {G : C} (h : IsCoseparator G) (α : C ≌ D) : + IsCoseparator (α.functor.obj G) := by simpa using IsCoseparating.of_equivalence h α + +end Equivalence + section Dual theorem isSeparator_op_iff (G : C) : IsSeparator (op G) ↔ IsCoseparator G := by @@ -704,6 +736,16 @@ instance HasSeparator.wellPowered [HasPullbacks C] [Balanced C] [HasSeparator C] end Instances +section Equivalence + +theorem HasSeparator.of_equivalence [HasSeparator C] (α : C ≌ D) : HasSeparator D := + ⟨α.functor.obj (separator C), isSeparator_separator C |>.of_equivalence α⟩ + +theorem HasCoseparator.of_equivalence [HasCoseparator C] (α : C ≌ D) : HasCoseparator D := + ⟨α.functor.obj (coseparator C), isCoseparator_coseparator C |>.of_equivalence α⟩ + +end Equivalence + section Dual @[simp] From 0f3e538b0b562f4ecf8baeebd1337c4c1ed138e7 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Mon, 16 Dec 2024 18:13:08 +0000 Subject: [PATCH 749/829] chore: move FreeLocus to AlgebraicGeometry/PrimeSpectrum (#20004) The imports of this file were relatively heavy: AlgebraicGeometry, RingTheory, Topology. So this place seems more appropriate. --- Mathlib.lean | 2 +- .../Module => AlgebraicGeometry/PrimeSpectrum}/FreeLocus.lean | 0 2 files changed, 1 insertion(+), 1 deletion(-) rename Mathlib/{Algebra/Module => AlgebraicGeometry/PrimeSpectrum}/FreeLocus.lean (100%) diff --git a/Mathlib.lean b/Mathlib.lean index d3490bf12958a..032b33531b73b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -528,7 +528,6 @@ import Mathlib.Algebra.Module.Equiv.Basic import Mathlib.Algebra.Module.Equiv.Defs import Mathlib.Algebra.Module.Equiv.Opposite import Mathlib.Algebra.Module.FinitePresentation -import Mathlib.Algebra.Module.FreeLocus import Mathlib.Algebra.Module.GradedModule import Mathlib.Algebra.Module.Hom import Mathlib.Algebra.Module.Injective @@ -994,6 +993,7 @@ import Mathlib.AlgebraicGeometry.Noetherian import Mathlib.AlgebraicGeometry.OpenImmersion import Mathlib.AlgebraicGeometry.Over import Mathlib.AlgebraicGeometry.PrimeSpectrum.Basic +import Mathlib.AlgebraicGeometry.PrimeSpectrum.FreeLocus import Mathlib.AlgebraicGeometry.PrimeSpectrum.IsOpenComapC import Mathlib.AlgebraicGeometry.PrimeSpectrum.Jacobson import Mathlib.AlgebraicGeometry.PrimeSpectrum.Maximal diff --git a/Mathlib/Algebra/Module/FreeLocus.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/FreeLocus.lean similarity index 100% rename from Mathlib/Algebra/Module/FreeLocus.lean rename to Mathlib/AlgebraicGeometry/PrimeSpectrum/FreeLocus.lean From 24a676818870a6a08c8c86324abf4e672c8e7958 Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Mon, 16 Dec 2024 22:01:34 +0000 Subject: [PATCH 750/829] feat(CategoryTheory): Finality of `StructuredArrow.post` (#20001) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit With a slight generalization to `map₂ (R' := T ⋙ S) (F := 𝟭 _) u (𝟙 (T ⋙ S))`. --- .../Comma/StructuredArrow/Basic.lean | 42 ++++++++++++++++--- Mathlib/CategoryTheory/Filtered/Final.lean | 34 ++++++++++++++- 2 files changed, 69 insertions(+), 7 deletions(-) diff --git a/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean b/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean index ec0b1cbac333c..22de055a81308 100644 --- a/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean +++ b/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean @@ -329,6 +329,11 @@ noncomputable instance isEquivalenceMap₂ end +/-- `StructuredArrow.post` is a special case of `StructuredArrow.map₂` up to natural isomorphism. -/ +def postIsoMap₂ (S : C) (F : B ⥤ C) (G : C ⥤ D) : + post S F G ≅ map₂ (F := 𝟭 _) (𝟙 _) (𝟙 (F ⋙ G)) := + NatIso.ofComponents fun _ => isoMk <| Iso.refl _ + /-- A structured arrow is called universal if it is initial. -/ abbrev IsUniversal (f : StructuredArrow S T) := IsInitial f @@ -665,6 +670,12 @@ noncomputable instance isEquivalenceMap₂ end +/-- `CostructuredArrow.post` is a special case of `CostructuredArrow.map₂` up to natural +isomorphism. -/ +def postIsoMap₂ (S : C) (F : B ⥤ C) (G : C ⥤ D) : + post F G S ≅ map₂ (F := 𝟭 _) (𝟙 (F ⋙ G)) (𝟙 _) := + NatIso.ofComponents fun _ => isoMk <| Iso.refl _ + /-- A costructured arrow is called universal if it is terminal. -/ abbrev IsUniversal (f : CostructuredArrow S T) := IsTerminal f @@ -875,7 +886,7 @@ variable {E : Type u₃} [Category.{v₃} E] (F : C ⥤ D) {G : D ⥤ E} {e : E} /-- The functor establishing the equivalence `StructuredArrow.preEquivalence`. -/ @[simps!] -def StructuredArrow.preEquivalence.functor (f : StructuredArrow e G) : +def StructuredArrow.preEquivalenceFunctor (f : StructuredArrow e G) : StructuredArrow f (pre e F G) ⥤ StructuredArrow f.right F where obj g := mk g.hom.right map φ := homMk φ.right.right <| by @@ -886,7 +897,7 @@ def StructuredArrow.preEquivalence.functor (f : StructuredArrow e G) : /-- The inverse functor establishing the equivalence `StructuredArrow.preEquivalence`. -/ @[simps!] -def StructuredArrow.preEquivalence.inverse (f : StructuredArrow e G) : +def StructuredArrow.preEquivalenceInverse (f : StructuredArrow e G) : StructuredArrow f.right F ⥤ StructuredArrow f (pre e F G) where obj g := mk (Y := mk (Y := g.right) @@ -898,13 +909,23 @@ def StructuredArrow.preEquivalence.inverse (f : StructuredArrow e G) : /-- A structured arrow category on a `StructuredArrow.pre e F G` functor is equivalent to the structured arrow category on F -/ +@[simps] def StructuredArrow.preEquivalence (f : StructuredArrow e G) : StructuredArrow f (pre e F G) ≌ StructuredArrow f.right F where - functor := StructuredArrow.preEquivalence.functor F f - inverse := StructuredArrow.preEquivalence.inverse F f + functor := preEquivalenceFunctor F f + inverse := preEquivalenceInverse F f unitIso := NatIso.ofComponents (fun _ => isoMk (isoMk (Iso.refl _))) counitIso := NatIso.ofComponents (fun _ => isoMk (Iso.refl _)) +/-- The functor `StructuredArrow d T ⥤ StructuredArrow e (T ⋙ S)` that `u : e ⟶ S.obj d` +induces via `StructuredArrow.map₂` can be expressed up to isomorphism by +`StructuredArrow.preEquivalence` and `StructuredArrow.proj`. -/ +def StructuredArrow.map₂IsoPreEquivalenceInverseCompProj (T : C ⥤ D) (S : D ⥤ E) (d : D) (e : E) + (u : e ⟶ S.obj d) : + map₂ (F := 𝟭 _) u (𝟙 (T ⋙ S)) ≅ + (preEquivalence T (mk u)).inverse ⋙ proj (mk u) (pre _ T S) := + NatIso.ofComponents fun _ => isoMk (Iso.refl _) + /-- The functor establishing the equivalence `CostructuredArrow.preEquivalence`. -/ @[simps!] def CostructuredArrow.preEquivalence.functor (f : CostructuredArrow G e) : @@ -929,11 +950,20 @@ def CostructuredArrow.preEquivalence.inverse (f : CostructuredArrow G e) : costructured arrow category on F -/ def CostructuredArrow.preEquivalence (f : CostructuredArrow G e) : CostructuredArrow (pre F G e) f ≌ CostructuredArrow F f.left where - functor := CostructuredArrow.preEquivalence.functor F f - inverse := CostructuredArrow.preEquivalence.inverse F f + functor := preEquivalence.functor F f + inverse := preEquivalence.inverse F f unitIso := NatIso.ofComponents (fun _ => isoMk (isoMk (Iso.refl _))) counitIso := NatIso.ofComponents (fun _ => isoMk (Iso.refl _)) +/-- The functor `CostructuredArrow T d ⥤ CostructuredArrow (T ⋙ S) e` that `u : S.obj d ⟶ e` +induces via `CostructuredArrow.map₂` can be expressed up to isomorphism by +`CostructuredArrow.preEquivalence` and `CostructuredArrow.proj`. -/ +def CostructuredArrow.map₂IsoPreEquivalenceInverseCompProj (T : C ⥤ D) (S : D ⥤ E) (d : D) (e : E) + (u : S.obj d ⟶ e) : + map₂ (F := 𝟭 _) (U := T ⋙ S) (𝟙 (T ⋙ S)) u ≅ + (preEquivalence T (mk u)).inverse ⋙ proj (pre T S _) (mk u) := + NatIso.ofComponents fun _ => isoMk (Iso.refl _) + end Pre end CategoryTheory diff --git a/Mathlib/CategoryTheory/Filtered/Final.lean b/Mathlib/CategoryTheory/Filtered/Final.lean index 55f7ae63b2a51..8710c29ec9105 100644 --- a/Mathlib/CategoryTheory/Filtered/Final.lean +++ b/Mathlib/CategoryTheory/Filtered/Final.lean @@ -25,6 +25,7 @@ final can be restated. We show: `C` is filtered and `F` is final. * Finality and initiality of diagonal functors `diag : C ⥤ C × C` and of projection functors of (co)structured arrow categories. +* Finality of `StructuredArrow.post`, given the finality of its arguments. ## References @@ -32,7 +33,7 @@ final can be restated. We show: -/ -universe v₁ v₂ u₁ u₂ +universe v₁ v₂ v₃ u₁ u₂ u₃ namespace CategoryTheory @@ -349,6 +350,37 @@ instance CostructuredArrow.initial_proj_of_isCofiltered [IsCofilteredOrEmpty C] rw [isConnected_iff_of_equivalence (ofCostructuredArrowProjEquivalence T Y X)] exact (initial_comp (Over.forget X) T).out _ +/-- The functor `StructuredArrow d T ⥤ StructuredArrow e (T ⋙ S)` that `u : e ⟶ S.obj d` +induces via `StructuredArrow.map₂` is final, if `T` and `S` are final and the domain of `T` is +filtered. -/ +instance StructuredArrow.final_map₂_id {C : Type v₁} [Category.{v₁} C] [IsFiltered C] {E : Type u₃} + [Category.{v₁} E] (T : C ⥤ D) [T.Final] (S : D ⥤ E) [S.Final] (d : D) (e : E) + (u : e ⟶ S.obj d) : Final (map₂ (R' := T ⋙ S) (F := 𝟭 _) u (𝟙 (T ⋙ S))) := by + have := (T ⋙ S).final_iff_isFiltered_structuredArrow.mp inferInstance e + apply final_of_natIso (map₂IsoPreEquivalenceInverseCompProj T S d e u).symm + +/-- `StructuredArrow.post X T S` is final if `T` and `S` are final and the domain of `T` is +filtered. -/ +instance StructuredArrow.final_post {C : Type v₁} [Category.{v₁} C] [IsFiltered C] {E : Type u₃} + [Category.{v₁} E] (X : D) (T : C ⥤ D) [T.Final] (S : D ⥤ E) [S.Final] : Final (post X T S) := by + apply final_of_natIso (postIsoMap₂ X T S).symm + +/-- The functor `CostructuredArrow T d ⥤ CostructuredArrow (T ⋙ S) e` that `u : S.obj d ⟶ e` +induces via `CostructuredArrow.map₂` is initial, if `T` and `S` are initial and the domain of `T` is +filtered. -/ +instance CostructuredArrow.initial_map₂_id {C : Type v₁} [Category.{v₁} C] [IsCofiltered C] + {E : Type u₃} [Category.{v₁} E] (T : C ⥤ D) [T.Initial] (S : D ⥤ E) [S.Initial] (d : D) (e : E) + (u : S.obj d ⟶ e) : Initial (map₂ (F := 𝟭 _) (U := T ⋙ S) (𝟙 (T ⋙ S)) u) := by + have := (T ⋙ S).initial_iff_isCofiltered_costructuredArrow.mp inferInstance e + apply initial_of_natIso (map₂IsoPreEquivalenceInverseCompProj T S d e u).symm + +/-- `CostructuredArrow.post T S X` is initial if `T` and `S` are initial and the domain of `T` is +cofiltered. -/ +instance CostructuredArrow.initial_post {C : Type v₁} [Category.{v₁} C] [IsCofiltered C] + {E : Type u₃} [Category.{v₁} E] (X : D) (T : C ⥤ D) [T.Initial] (S : D ⥤ E) [S.Initial] : + Initial (post T S X) := by + apply initial_of_natIso (postIsoMap₂ X T S).symm + section Pi variable {α : Type u₁} {I : α → Type u₂} [∀ s, Category.{v₂} (I s)] From 4ea2e9180d14ae30564306b3e3e69f3039c888f3 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Mon, 16 Dec 2024 22:32:24 +0000 Subject: [PATCH 751/829] feat: add `IsLeast {n | p n} (Nat.find hp)` and related lemmas (#20010) --- Mathlib/Data/Int/LeastGreatest.lean | 13 +++++++++++++ Mathlib/Order/Nat.lean | 12 ++++++++++++ 2 files changed, 25 insertions(+) diff --git a/Mathlib/Data/Int/LeastGreatest.lean b/Mathlib/Data/Int/LeastGreatest.lean index 9e3c23e0cbff1..d5d12b857a00a 100644 --- a/Mathlib/Data/Int/LeastGreatest.lean +++ b/Mathlib/Data/Int/LeastGreatest.lean @@ -5,6 +5,7 @@ Authors: Jeremy Avigad, Mario Carneiro -/ import Mathlib.Algebra.Order.Ring.Int import Mathlib.Data.Nat.Find +import Mathlib.Order.Bounds.Defs /-! # Least upper bound and greatest lower bound properties for integers @@ -51,6 +52,11 @@ def leastOfBdd {P : ℤ → Prop} [DecidablePred P] (b : ℤ) (Hb : ∀ z : ℤ, match z, le.dest (Hb _ h), h with | _, ⟨_, rfl⟩, h => add_le_add_left (Int.ofNat_le.2 <| Nat.find_min' _ h) _⟩ +/-- `Int.leastOfBdd` is the least integer satisfying a predicate which is false for all `z : ℤ` with +`z < b` for some fixed `b : ℤ`. -/ +lemma isLeast_coe_leastOfBdd {P : ℤ → Prop} [DecidablePred P] (b : ℤ) (Hb : ∀ z : ℤ, P z → b ≤ z) + (Hinh : ∃ z : ℤ, P z) : IsLeast {z | P z} (leastOfBdd b Hb Hinh : ℤ) := + (leastOfBdd b Hb Hinh).2 /-- If `P : ℤ → Prop` is a predicate such that the set `{m : P m}` is bounded below and nonempty, @@ -84,6 +90,13 @@ def greatestOfBdd {P : ℤ → Prop} [DecidablePred P] (b : ℤ) (Hb : ∀ z : let ⟨lb, Plb, al⟩ := leastOfBdd (-b) Hbdd' Hinh' ⟨-lb, Plb, fun z h => le_neg.1 <| al _ <| by rwa [neg_neg]⟩ +/-- `Int.greatestOfBdd` is the greatest integer satisfying a predicate which is false for all +`z : ℤ` with `b < z` for some fixed `b : ℤ`. -/ +lemma isGreatest_coe_greatestOfBdd {P : ℤ → Prop} [DecidablePred P] (b : ℤ) + (Hb : ∀ z : ℤ, P z → z ≤ b) (Hinh : ∃ z : ℤ, P z) : + IsGreatest {z | P z} (greatestOfBdd b Hb Hinh : ℤ) := + (greatestOfBdd b Hb Hinh).2 + /-- If `P : ℤ → Prop` is a predicate such that the set `{m : P m}` is bounded above and nonempty, then this set has the greatest element. This lemma uses classical logic to avoid assumption diff --git a/Mathlib/Order/Nat.lean b/Mathlib/Order/Nat.lean index 857c598eeebea..44594c91b5e4a 100644 --- a/Mathlib/Order/Nat.lean +++ b/Mathlib/Order/Nat.lean @@ -3,7 +3,9 @@ Copyright (c) 2014 Floris van Doorn (c) 2016 Microsoft Corporation. All rights r Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ +import Mathlib.Data.Nat.Find import Mathlib.Order.BoundedOrder.Basic +import Mathlib.Order.Bounds.Defs /-! # The natural numbers form a linear order @@ -31,4 +33,14 @@ instance instNoMaxOrder : NoMaxOrder ℕ where -- We want to use this lemma earlier than the lemma simp can prove it with @[simp, nolint simpNF] protected lemma bot_eq_zero : ⊥ = 0 := rfl +/-- `Nat.find` is the minimum natural number satisfying a predicate `p`. -/ +lemma isLeast_find {p : ℕ → Prop} [DecidablePred p] (hp : ∃ n, p n) : + IsLeast {n | p n} (Nat.find hp) := + ⟨Nat.find_spec hp, fun _ ↦ Nat.find_min' hp⟩ + end Nat + +/-- `Nat.find` is the minimum element of a nonempty set of natural numbers. -/ +lemma Set.Nonempty.isLeast_natFind {s : Set ℕ} [DecidablePred (· ∈ s)] (hs : s.Nonempty) : + IsLeast s (Nat.find hs) := + Nat.isLeast_find hs From 458cc7d8784d3b2c54e3e142ae76e5aca576bc94 Mon Sep 17 00:00:00 2001 From: Noah Singer Date: Mon, 16 Dec 2024 22:48:21 +0000 Subject: [PATCH 752/829] chore: make `Decidable` instance for `SimpleGraph.Connected` be reducible (#20015) Suggested by @kmill. [![Open in Gitpod](https://gitpod.io/button/open-in-gitpod.svg)](https://gitpod.io/from-referrer/) --- .../SimpleGraph/Connectivity/WalkCounting.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean b/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean index 559e58c4930b6..01d7e6a4232f3 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean @@ -206,9 +206,9 @@ instance : Fintype G.ConnectedComponent := instance : Decidable G.Preconnected := inferInstanceAs <| Decidable (∀ u v, G.Reachable u v) -instance : Decidable G.Connected := by - rw [connected_iff, ← Finset.univ_nonempty_iff] - infer_instance +instance : Decidable G.Connected := + decidable_of_iff (G.Preconnected ∧ (Finset.univ : Finset V).Nonempty) <| by + rw [connected_iff, ← Finset.univ_nonempty_iff] instance Path.instFintype {u v : V} : Fintype (G.Path u v) where elems := (univ (α := { p : G.Walk u v | p.IsPath ∧ p.length < Fintype.card V })).map From b18fc8a82981a8fbf1c48387e28cd63422e3fccb Mon Sep 17 00:00:00 2001 From: Yoh Tanimoto Date: Tue, 17 Dec 2024 01:35:05 +0000 Subject: [PATCH 753/829] feat(MeasureTheory/Integral/RieszMarkovKakutani) prove that the Riesz content is indeed a content for `NNReal` (#18775) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Prove additivity of `RieszContentAux`, completing the proof that it gives a `Content`. Motivation: this gives the measure related with a positive linear functional `Λ` on compactly supported continuous functions. The next step is to characterise the constructed measure as the one giving `Λ` back. In this PR, it is assumed to be `NNReal`-linear. The main steps to prove additivity have been proposed by @kkytola [here](https://github.com/leanprover-community/mathlib4/pull/12290#issuecomment-2125859907). A different approach is taken in #12290 for `Real`-linear `Λ`. - [x] depends on: #12266 for a variation of `PartitionOfUnity`. The new contents of this PR are contained in `MeasureTheory.Integral.linearRieszMarkovKakutani` and `Topology.ContinuousMap.CompactlySupported`. Co-authored-by: Yoh Tanimoto <57562556+yoh-tanimoto@users.noreply.github.com> --- .../Integral/RieszMarkovKakutani.lean | 206 +++++++++++++++--- .../ContinuousMap/CompactlySupported.lean | 22 ++ 2 files changed, 198 insertions(+), 30 deletions(-) diff --git a/Mathlib/MeasureTheory/Integral/RieszMarkovKakutani.lean b/Mathlib/MeasureTheory/Integral/RieszMarkovKakutani.lean index 0c8f7657a9261..5e34c8f3090f7 100644 --- a/Mathlib/MeasureTheory/Integral/RieszMarkovKakutani.lean +++ b/Mathlib/MeasureTheory/Integral/RieszMarkovKakutani.lean @@ -3,18 +3,20 @@ 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.Sets.Compacts -import Mathlib.Topology.ContinuousMap.Bounded.Basic +import Mathlib.MeasureTheory.Measure.Content +import Mathlib.Topology.ContinuousMap.CompactlySupported +import Mathlib.Topology.PartitionOfUnity /-! # Riesz–Markov–Kakutani representation theorem -This file will prove different versions of the Riesz-Markov-Kakutani representation theorem. -The theorem is first proven for compact spaces, from which the statements about linear functionals -on bounded continuous functions or compactly supported functions on locally compact spaces follow. +This file will prove the Riesz-Markov-Kakutani representation theorem on a locally compact +T2 space `X`. As a special case, the statements about linear functionals on bounded continuous +functions follows. To make use of the existing API, the measure is constructed from a content `λ` on the -compact subsets of the space X, rather than the usual construction of open sets in the literature. +compact subsets of a locally compact space X, rather than the usual construction of open sets in the +literature. ## References @@ -25,35 +27,51 @@ compact subsets of the space X, rather than the usual construction of open sets noncomputable section -open BoundedContinuousFunction NNReal ENNReal - -open Set Function TopologicalSpace +open scoped BoundedContinuousFunction NNReal ENNReal +open Set Function TopologicalSpace CompactlySupported CompactlySupportedContinuousMap + MeasureTheory variable {X : Type*} [TopologicalSpace X] -variable (Λ : (X →ᵇ ℝ≥0) →ₗ[ℝ≥0] ℝ≥0) +variable (Λ : C_c(X, ℝ≥0) →ₗ[ℝ≥0] ℝ≥0) /-! ### Construction of the content: -/ - -/-- Given a positive linear functional Λ on X, for `K ⊆ X` compact define -`λ(K) = inf {Λf | 1≤f on K}`. When X is a compact Hausdorff space, this will be shown to be a +/-- Given a positive linear functional `Λ` on continuous compactly supported functions on `X` +with values in `ℝ≥0`, for `K ⊆ X` compact define `λ(K) = inf {Λf | 1≤f on K}`. +When `X` is a locally compact T2 space, this will be shown to be a content, and will be shown to agree with the Riesz measure on the compact subsets `K ⊆ X`. -/ def rieszContentAux : Compacts X → ℝ≥0 := fun K => - sInf (Λ '' { f : X →ᵇ ℝ≥0 | ∀ x ∈ K, (1 : ℝ≥0) ≤ f x }) + sInf (Λ '' { f : C_c(X, ℝ≥0) | ∀ x ∈ K, (1 : ℝ≥0) ≤ f x }) section RieszMonotone -/-- For any compact subset `K ⊆ X`, there exist some bounded continuous nonnegative -functions f on X such that `f ≥ 1` on K. -/ +variable [T2Space X] [LocallyCompactSpace X] + +/-- For any compact subset `K ⊆ X`, there exist some compactly supported continuous nonnegative +functions `f` on `X` such that `f ≥ 1` on `K`. -/ theorem rieszContentAux_image_nonempty (K : Compacts X) : - (Λ '' { f : X →ᵇ ℝ≥0 | ∀ x ∈ K, (1 : ℝ≥0) ≤ f x }).Nonempty := by + (Λ '' { f : C_c(X, ℝ≥0) | ∀ x ∈ K, (1 : ℝ≥0) ≤ f x }).Nonempty := by rw [image_nonempty] - use (1 : X →ᵇ ℝ≥0) - intro x _ - simp only [BoundedContinuousFunction.coe_one, Pi.one_apply]; rfl - -/-- Riesz content λ (associated with a positive linear functional Λ) is -monotone: if `K₁ ⊆ K₂` are compact subsets in X, then `λ(K₁) ≤ λ(K₂)`. -/ + obtain ⟨V, hVcp, hKsubintV⟩ := exists_compact_superset K.2 + have hIsCompact_closure_interior : IsCompact (closure (interior V)) := by + apply IsCompact.of_isClosed_subset hVcp isClosed_closure + nth_rw 2 [← closure_eq_iff_isClosed.mpr (IsCompact.isClosed hVcp)] + exact closure_mono interior_subset + obtain ⟨f, hsuppfsubV, hfeq1onK, hfinicc⟩ := + exists_tsupport_one_of_isOpen_isClosed isOpen_interior hIsCompact_closure_interior + (IsCompact.isClosed K.2) hKsubintV + have hfHasCompactSupport : HasCompactSupport f := + IsCompact.of_isClosed_subset hVcp (isClosed_tsupport f) + (Set.Subset.trans hsuppfsubV interior_subset) + use nnrealPart ⟨f, hfHasCompactSupport⟩ + intro x hx + apply le_of_eq + simp only [nnrealPart_apply, CompactlySupportedContinuousMap.coe_mk] + rw [← Real.toNNReal_one, Real.toNNReal_eq_toNNReal_iff (zero_le_one' ℝ) (hfinicc x).1] + exact hfeq1onK.symm hx + +/-- Riesz content `λ` (associated with a positive linear functional `Λ`) is +monotone: if `K₁ ⊆ K₂` are compact subsets in `X`, then `λ(K₁) ≤ λ(K₂)`. -/ theorem rieszContentAux_mono {K₁ K₂ : Compacts X} (h : K₁ ≤ K₂) : rieszContentAux Λ K₁ ≤ rieszContentAux Λ K₂ := csInf_le_csInf (OrderBot.bddBelow _) (rieszContentAux_image_nonempty Λ K₂) @@ -63,17 +81,19 @@ end RieszMonotone section RieszSubadditive -/-- Any bounded continuous nonnegative f such that `f ≥ 1` on K gives an upper bound on the -content of K; namely `λ(K) ≤ Λ f`. -/ -theorem rieszContentAux_le {K : Compacts X} {f : X →ᵇ ℝ≥0} (h : ∀ x ∈ K, (1 : ℝ≥0) ≤ f x) : +/-- Any compactly supported continuous nonnegative `f` such that `f ≥ 1` on `K` gives an upper bound +on the content of `K`; namely `λ(K) ≤ Λ f`. -/ +theorem rieszContentAux_le {K : Compacts X} {f : C_c(X, ℝ≥0)} (h : ∀ x ∈ K, (1 : ℝ≥0) ≤ f x) : rieszContentAux Λ K ≤ Λ f := csInf_le (OrderBot.bddBelow _) ⟨f, ⟨h, rfl⟩⟩ +variable [T2Space X] [LocallyCompactSpace X] + /-- The Riesz content can be approximated arbitrarily well by evaluating the positive linear -functional on test functions: for any `ε > 0`, there exists a bounded continuous nonnegative -function f on X such that `f ≥ 1` on K and such that `λ(K) ≤ Λ f < λ(K) + ε`. -/ +functional on test functions: for any `ε > 0`, there exists a compactly supported continuous +nonnegative function `f` on `X` such that `f ≥ 1` on `K` and such that `λ(K) ≤ Λ f < λ(K) + ε`. -/ theorem exists_lt_rieszContentAux_add_pos (K : Compacts X) {ε : ℝ≥0} (εpos : 0 < ε) : - ∃ f : X →ᵇ ℝ≥0, (∀ x ∈ K, (1 : ℝ≥0) ≤ f x) ∧ Λ f < rieszContentAux Λ K + ε := by + ∃ f : C_c(X, ℝ≥0), (∀ x ∈ K, (1 : ℝ≥0) ≤ f x) ∧ Λ f < rieszContentAux Λ K + ε := by --choose a test function `f` s.t. `Λf = α < λ(K) + ε` obtain ⟨α, ⟨⟨f, f_hyp⟩, α_hyp⟩⟩ := exists_lt_of_csInf_lt (rieszContentAux_image_nonempty Λ K) @@ -82,7 +102,7 @@ theorem exists_lt_rieszContentAux_add_pos (K : Compacts X) {ε : ℝ≥0} (εpos rw [f_hyp.right] exact α_hyp -/-- The Riesz content λ associated to a given positive linear functional Λ is +/-- The Riesz content `λ` associated to a given positive linear functional `Λ` is finitely subadditive: `λ(K₁ ∪ K₂) ≤ λ(K₁) + λ(K₂)` for any compact subsets `K₁, K₂ ⊆ X`. -/ theorem rieszContentAux_sup_le (K1 K2 : Compacts X) : rieszContentAux Λ (K1 ⊔ K2) ≤ rieszContentAux Λ K1 + rieszContentAux Λ K2 := by @@ -105,3 +125,129 @@ theorem rieszContentAux_sup_le (K1 K2 : Compacts X) : rw [add_assoc, add_comm (ε / 2), add_assoc, add_halves ε, add_assoc] end RieszSubadditive + + +section PartitionOfUnity + +variable [T2Space X] [LocallyCompactSpace X] + +lemma exists_continuous_add_one_of_isCompact_nnreal + {s₀ s₁ : Set X} {t : Set X} (s₀_compact : IsCompact s₀) (s₁_compact : IsCompact s₁) + (t_compact : IsCompact t) (disj : Disjoint s₀ s₁) (hst : s₀ ∪ s₁ ⊆ t) : + ∃ (f₀ f₁ : C_c(X, ℝ≥0)), EqOn f₀ 1 s₀ ∧ EqOn f₁ 1 s₁ ∧ EqOn (f₀ + f₁) 1 t := by + set so : Fin 2 → Set X := fun j => if j = 0 then s₀ᶜ else s₁ᶜ with hso + have soopen (j : Fin 2) : IsOpen (so j) := by + fin_cases j + · simp only [hso, Fin.zero_eta, Fin.isValue, ↓reduceIte, isOpen_compl_iff] + exact IsCompact.isClosed <| s₀_compact + · simp only [hso, Fin.isValue, Fin.mk_one, one_ne_zero, ↓reduceIte, isOpen_compl_iff] + exact IsCompact.isClosed <| s₁_compact + have hsot : t ⊆ ⋃ j, so j := by + rw [hso] + simp only [Fin.isValue] + intro x hx + rw [mem_iUnion] + rw [← subset_compl_iff_disjoint_right, ← compl_compl s₀, compl_subset_iff_union] at disj + have h : x ∈ s₀ᶜ ∨ x ∈ s₁ᶜ := by + rw [← mem_union, disj] + trivial + apply Or.elim h + · intro h0 + use 0 + simp only [Fin.isValue, ↓reduceIte] + exact h0 + · intro h1 + use 1 + simp only [Fin.isValue, one_ne_zero, ↓reduceIte] + exact h1 + obtain ⟨f, f_supp_in_so, sum_f_one_on_t, f_in_icc, f_hcs⟩ := + exists_continuous_sum_one_of_isOpen_isCompact soopen t_compact hsot + use (nnrealPart (⟨f 1, f_hcs 1⟩ : C_c(X, ℝ))), + (nnrealPart (⟨f 0, f_hcs 0⟩ : C_c(X, ℝ))) + simp only [Fin.isValue, CompactlySupportedContinuousMap.coe_add] + have sum_one_x (x : X) (hx : x ∈ t) : (f 0) x + (f 1) x = 1 := by + simpa only [Finset.sum_apply, Fin.sum_univ_two, Fin.isValue, Pi.one_apply] + using sum_f_one_on_t hx + refine ⟨?_, ?_, ?_⟩ + · intro x hx + simp only [Fin.isValue, nnrealPart_apply, + CompactlySupportedContinuousMap.coe_mk, Pi.one_apply, Real.toNNReal_eq_one] + have : (f 0) x = 0 := by + rw [← nmem_support] + have : s₀ ⊆ (tsupport (f 0))ᶜ := by + apply subset_trans _ (compl_subset_compl.mpr (f_supp_in_so 0)) + rw [hso] + simp only [Fin.isValue, ↓reduceIte, compl_compl, subset_refl] + apply not_mem_of_mem_compl + exact mem_of_subset_of_mem (subset_trans this (compl_subset_compl_of_subset subset_closure)) + hx + rw [union_subset_iff] at hst + rw [← sum_one_x x (mem_of_subset_of_mem hst.1 hx), this] + exact Eq.symm (AddZeroClass.zero_add ((f 1) x)) + · intro x hx + simp only [Fin.isValue, nnrealPart_apply, + CompactlySupportedContinuousMap.coe_mk, Pi.one_apply, Real.toNNReal_eq_one] + have : (f 1) x = 0 := by + rw [← nmem_support] + have : s₁ ⊆ (tsupport (f 1))ᶜ := by + apply subset_trans _ (compl_subset_compl.mpr (f_supp_in_so 1)) + rw [hso] + simp only [Fin.isValue, one_ne_zero, ↓reduceIte, compl_compl, subset_refl] + apply not_mem_of_mem_compl + exact mem_of_subset_of_mem (subset_trans this (compl_subset_compl_of_subset subset_closure)) + hx + rw [union_subset_iff] at hst + rw [← sum_one_x x (mem_of_subset_of_mem hst.2 hx), this] + exact Eq.symm (AddMonoid.add_zero ((f 0) x)) + · intro x hx + simp only [Fin.isValue, Pi.add_apply, nnrealPart_apply, + CompactlySupportedContinuousMap.coe_mk, Pi.one_apply] + rw [Real.toNNReal_add_toNNReal (f_in_icc 1 x).1 (f_in_icc 0 x).1, add_comm] + simp only [Fin.isValue, Real.toNNReal_eq_one] + exact sum_one_x x hx + +end PartitionOfUnity + +variable [T2Space X] [LocallyCompactSpace X] + +lemma rieszContentAux_union {K₁ K₂ : TopologicalSpace.Compacts X} + (disj : Disjoint (K₁ : Set X) K₂) : + rieszContentAux Λ (K₁ ⊔ K₂) = rieszContentAux Λ K₁ + rieszContentAux Λ K₂ := by + refine le_antisymm (rieszContentAux_sup_le Λ K₁ K₂) ?_ + refine le_csInf (rieszContentAux_image_nonempty Λ (K₁ ⊔ K₂)) ?_ + intro b ⟨f, ⟨hf, Λf_eq_b⟩⟩ + have hsuppf : ∀ x ∈ K₁ ⊔ K₂, x ∈ support f := by + intro x hx + rw [mem_support] + exact Ne.symm (ne_of_lt <| lt_of_lt_of_le (zero_lt_one' ℝ≥0) (hf x hx)) + have hsubsuppf : (K₁ : Set X) ∪ (K₂ : Set X) ⊆ tsupport f := subset_trans hsuppf subset_closure + obtain ⟨g₁, g₂, hg₁, hg₂, sum_g⟩ := exists_continuous_add_one_of_isCompact_nnreal K₁.isCompact' + K₂.isCompact' f.hasCompactSupport'.isCompact disj hsubsuppf + have f_eq_sum : f = g₁ * f + g₂ * f := by + ext x + simp only [CompactlySupportedContinuousMap.coe_add, CompactlySupportedContinuousMap.coe_mul, + Pi.add_apply, Pi.mul_apply, NNReal.coe_add, NNReal.coe_mul, + Eq.symm (RightDistribClass.right_distrib _ _ _), ← NNReal.coe_add, ← Pi.add_apply] + by_cases h : f x = 0 + · rw [h] + simp only [NNReal.coe_zero, NNReal.coe_add, mul_zero] + · push_neg at h + simp only [CompactlySupportedContinuousMap.coe_add, ContinuousMap.toFun_eq_coe, + CompactlySupportedContinuousMap.coe_toContinuousMap] at sum_g + rw [sum_g (mem_of_subset_of_mem subset_closure (mem_support.mpr h))] + simp only [Pi.one_apply, NNReal.coe_one, one_mul] + rw [← Λf_eq_b, f_eq_sum, map_add] + have aux₁ : ∀ x ∈ K₁, 1 ≤ (g₁ * f) x := by + intro x x_in_K₁ + simp [hg₁ x_in_K₁, hf x (mem_union_left _ x_in_K₁)] + have aux₂ : ∀ x ∈ K₂, 1 ≤ (g₂ * f) x := by + intro x x_in_K₂ + simp [hg₂ x_in_K₂, hf x (mem_union_right _ x_in_K₂)] + exact add_le_add (rieszContentAux_le Λ aux₁) (rieszContentAux_le Λ aux₂) + +/-- The content induced by the linear functional `Λ`. -/ +noncomputable def rieszContent (Λ : C_c(X, ℝ≥0) →ₗ[ℝ≥0] ℝ≥0) : Content X where + toFun := rieszContentAux Λ + mono' := fun _ _ ↦ rieszContentAux_mono Λ + sup_disjoint' := fun _ _ disj _ _ ↦ rieszContentAux_union Λ disj + sup_le' := rieszContentAux_sup_le Λ diff --git a/Mathlib/Topology/ContinuousMap/CompactlySupported.lean b/Mathlib/Topology/ContinuousMap/CompactlySupported.lean index eb5348c9f448c..b57c2fb017f07 100644 --- a/Mathlib/Topology/ContinuousMap/CompactlySupported.lean +++ b/Mathlib/Topology/ContinuousMap/CompactlySupported.lean @@ -522,3 +522,25 @@ instance : ZeroAtInftyContinuousMapClass F β γ where end ZeroAtInfty end CompactlySupportedContinuousMapClass + +section NonnegativePart + +open NNReal + +namespace CompactlySupportedContinuousMap + +/-- The nonnegative part of a bounded continuous `ℝ`-valued function as a bounded +continuous `ℝ≥0`-valued function. -/ +noncomputable def nnrealPart (f : C_c(α, ℝ)) : C_c(α, ℝ≥0) where + toFun := Real.toNNReal.comp f.toFun + continuous_toFun := Continuous.comp continuous_real_toNNReal f.continuous + hasCompactSupport' := by + apply HasCompactSupport.comp_left f.hasCompactSupport' Real.toNNReal_zero + +@[simp] +lemma nnrealPart_apply (f : C_c(α, ℝ)) (x : α) : + f.nnrealPart x = Real.toNNReal (f x) := rfl + +end CompactlySupportedContinuousMap + +end NonnegativePart From d5bed454ee2c347b2d60a12fc609ed827c070617 Mon Sep 17 00:00:00 2001 From: "Filippo A. E. Nuccio" Date: Tue, 17 Dec 2024 11:38:20 +0000 Subject: [PATCH 754/829] faet(RingTheory/DiscreteValuationRing/Basic): rename `DiscreteValuationRing` to `IsDiscreteValuationRing`. (#19947) In this PR we rename `DiscreteValuationRing` to `IsDiscreteValuationRing`, since the latter is `Prop`-valued. --- .../NumberTheory/Padics/PadicIntegers.lean | 6 ++-- Mathlib/NumberTheory/Padics/RingHoms.lean | 2 +- Mathlib/RingTheory/DedekindDomain/Dvr.lean | 18 +++++----- Mathlib/RingTheory/DedekindDomain/PID.lean | 4 +-- .../DiscreteValuationRing/Basic.lean | 33 ++++++++++--------- .../DiscreteValuationRing/TFAE.lean | 16 ++++----- Mathlib/RingTheory/PowerSeries/Inverse.lean | 8 ++--- .../RingTheory/Valuation/ValuationRing.lean | 2 +- .../WittVector/DiscreteValuationRing.lean | 6 ++-- 9 files changed, 48 insertions(+), 47 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/PadicIntegers.lean b/Mathlib/NumberTheory/Padics/PadicIntegers.lean index 33bdb2db15211..163ab12838048 100644 --- a/Mathlib/NumberTheory/Padics/PadicIntegers.lean +++ b/Mathlib/NumberTheory/Padics/PadicIntegers.lean @@ -543,14 +543,14 @@ theorem prime_p : Prime (p : ℤ_[p]) := by theorem irreducible_p : Irreducible (p : ℤ_[p]) := Prime.irreducible prime_p -instance : DiscreteValuationRing ℤ_[p] := - DiscreteValuationRing.ofHasUnitMulPowIrreducibleFactorization +instance : IsDiscreteValuationRing ℤ_[p] := + IsDiscreteValuationRing.ofHasUnitMulPowIrreducibleFactorization ⟨p, irreducible_p, fun {x hx} => ⟨x.valuation.natAbs, unitCoeff hx, by rw [mul_comm, ← unitCoeff_spec hx]⟩⟩ theorem ideal_eq_span_pow_p {s : Ideal ℤ_[p]} (hs : s ≠ ⊥) : ∃ n : ℕ, s = Ideal.span {(p : ℤ_[p]) ^ n} := - DiscreteValuationRing.ideal_eq_span_pow_irreducible hs irreducible_p + IsDiscreteValuationRing.ideal_eq_span_pow_irreducible hs irreducible_p open CauSeq diff --git a/Mathlib/NumberTheory/Padics/RingHoms.lean b/Mathlib/NumberTheory/Padics/RingHoms.lean index 106a4b2f15a8e..8a7cb7bf0db3f 100644 --- a/Mathlib/NumberTheory/Padics/RingHoms.lean +++ b/Mathlib/NumberTheory/Padics/RingHoms.lean @@ -374,7 +374,7 @@ theorem appr_spec (n : ℕ) : ∀ x : ℤ_[p], x - appr x n ∈ Ideal.span {(p : -- TODO: write a `CanLift` instance for units rw [Int.natAbs_eq_zero] at hc0 rw [isUnit_iff, norm_eq_pow_val hc', hc0, neg_zero, zpow_zero] - rw [DiscreteValuationRing.unit_mul_pow_congr_unit _ _ _ _ _ hc] + rw [IsDiscreteValuationRing.unit_mul_pow_congr_unit _ _ _ _ _ hc] exact irreducible_p · simp only [zero_pow hc0, sub_zero, ZMod.cast_zero, mul_zero] rw [unitCoeff_spec hc'] diff --git a/Mathlib/RingTheory/DedekindDomain/Dvr.lean b/Mathlib/RingTheory/DedekindDomain/Dvr.lean index 5a719e51737c7..e927b6a209e68 100644 --- a/Mathlib/RingTheory/DedekindDomain/Dvr.lean +++ b/Mathlib/RingTheory/DedekindDomain/Dvr.lean @@ -54,7 +54,7 @@ This is equivalent to `IsDedekindDomain`. -/ class IsDedekindDomainDvr extends IsNoetherian A A : Prop where is_dvr_at_nonzero_prime : ∀ P ≠ (⊥ : Ideal A), ∀ _ : P.IsPrime, - DiscreteValuationRing (Localization.AtPrime P) + IsDiscreteValuationRing (Localization.AtPrime P) /-- Localizing a domain of Krull dimension `≤ 1` gives another ring of Krull dimension `≤ 1`. @@ -124,16 +124,16 @@ theorem IsLocalization.AtPrime.not_isField {P : Ideal A} (hP : P ≠ ⊥) [pP : x_ne))) /-- In a Dedekind domain, the localization at every nonzero prime ideal is a DVR. -/ -theorem IsLocalization.AtPrime.discreteValuationRing_of_dedekind_domain [IsDedekindDomain A] +theorem IsLocalization.AtPrime.isDiscreteValuationRing_of_dedekind_domain [IsDedekindDomain A] {P : Ideal A} (hP : P ≠ ⊥) [pP : P.IsPrime] (Aₘ : Type*) [CommRing Aₘ] [IsDomain Aₘ] - [Algebra A Aₘ] [IsLocalization.AtPrime Aₘ P] : DiscreteValuationRing Aₘ := by + [Algebra A Aₘ] [IsLocalization.AtPrime Aₘ P] : IsDiscreteValuationRing Aₘ := by classical letI : IsNoetherianRing Aₘ := IsLocalization.isNoetherianRing P.primeCompl _ IsDedekindRing.toIsNoetherian letI : IsLocalRing Aₘ := IsLocalization.AtPrime.isLocalRing Aₘ P have hnf := IsLocalization.AtPrime.not_isField A hP Aₘ exact - ((DiscreteValuationRing.TFAE Aₘ hnf).out 0 2).mpr + ((IsDiscreteValuationRing.TFAE Aₘ hnf).out 0 2).mpr (IsLocalization.AtPrime.isDedekindDomain A P _) /-- Dedekind domains, in the sense of Noetherian integrally closed domains of Krull dimension ≤ 1, @@ -141,7 +141,7 @@ are also Dedekind domains in the sense of Noetherian domains where the localizat nonzero prime ideal is a DVR. -/ instance IsDedekindDomain.isDedekindDomainDvr [IsDedekindDomain A] : IsDedekindDomainDvr A where is_dvr_at_nonzero_prime := fun _ hP _ => - IsLocalization.AtPrime.discreteValuationRing_of_dedekind_domain A hP _ + IsLocalization.AtPrime.isDiscreteValuationRing_of_dedekind_domain A hP _ instance IsDedekindDomainDvr.ring_dimensionLEOne [h : IsDedekindDomainDvr A] : Ring.DimensionLEOne A where @@ -156,7 +156,7 @@ instance IsDedekindDomainDvr.ring_dimensionLEOne [h : IsDedekindDomainDvr A] : have hp1 : P.1 ≠ ⊥ := fun x => hp ((p.map_eq_bot_iff_of_injective hinj).mp x) have hq1 : Q.1 ≠ ⊥ := fun x => (ne_bot_of_le_ne_bot hp hpq) ((q.map_eq_bot_iff_of_injective hinj).mp x) - rcases (DiscreteValuationRing.iff_pid_with_one_nonzero_prime (Localization.AtPrime q)).mp + rcases (IsDiscreteValuationRing.iff_pid_with_one_nonzero_prime (Localization.AtPrime q)).mp (h.is_dvr_at_nonzero_prime q (ne_bot_of_le_ne_bot hp hpq) hq.isPrime) with ⟨_, huq⟩ rw [show p = q from Subtype.val_inj.mpr <| f.injective <| Subtype.val_inj.mp (huq.unique ⟨hp1, P.2⟩ ⟨hq1, Q.2⟩)] @@ -164,9 +164,9 @@ instance IsDedekindDomainDvr.ring_dimensionLEOne [h : IsDedekindDomainDvr A] : instance IsDedekindDomainDvr.isIntegrallyClosed [h : IsDedekindDomainDvr A] : IsIntegrallyClosed A := - IsIntegrallyClosed.of_localization_maximal <| fun p hp0 hpm => - let ⟨_, _⟩ := (DiscreteValuationRing.iff_pid_with_one_nonzero_prime (Localization.AtPrime p)).mp - (h.is_dvr_at_nonzero_prime p hp0 hpm.isPrime) + IsIntegrallyClosed.of_localization_maximal <| fun p hp0 hpm ↦ + let ⟨_, _⟩ := (IsDiscreteValuationRing.iff_pid_with_one_nonzero_prime + (Localization.AtPrime p)).mp (h.is_dvr_at_nonzero_prime p hp0 hpm.isPrime) inferInstance /-- If an integral domain is Noetherian, and the localization at every nonzero prime is diff --git a/Mathlib/RingTheory/DedekindDomain/PID.lean b/Mathlib/RingTheory/DedekindDomain/PID.lean index 673e136d6943e..b9bad1bafa86b 100644 --- a/Mathlib/RingTheory/DedekindDomain/PID.lean +++ b/Mathlib/RingTheory/DedekindDomain/PID.lean @@ -210,8 +210,8 @@ theorem IsLocalization.OverPrime.mem_normalizedFactors_of_isPrime [IsDomain S] exact (IsLocalization.map_eq (T := Algebra.algebraMapSubmonoid S (primeCompl p)) (Submonoid.le_comap_map _) x).symm obtain ⟨pid, p', ⟨hp'0, hp'p⟩, hpu⟩ := - (DiscreteValuationRing.iff_pid_with_one_nonzero_prime (Localization.AtPrime p)).mp - (IsLocalization.AtPrime.discreteValuationRing_of_dedekind_domain R hp0 _) + (IsDiscreteValuationRing.iff_pid_with_one_nonzero_prime (Localization.AtPrime p)).mp + (IsLocalization.AtPrime.isDiscreteValuationRing_of_dedekind_domain R hp0 _) have : IsLocalRing.maximalIdeal (Localization.AtPrime p) ≠ ⊥ := by rw [Submodule.ne_bot_iff] at hp0 ⊢ obtain ⟨x, x_mem, x_ne⟩ := hp0 diff --git a/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean b/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean index 33a536b988ff1..379e1e61eb9bf 100644 --- a/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean +++ b/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean @@ -24,7 +24,7 @@ book "Local Fields"). Let R be an integral domain, assumed to be a principal ideal ring and a local ring. -* `DiscreteValuationRing R` : a predicate expressing that R is a DVR. +* `IsDiscreteValuationRing R` : a predicate expressing that R is a DVR. ### Definitions @@ -46,13 +46,13 @@ open Ideal IsLocalRing /-- An integral domain is a *discrete valuation ring* (DVR) if it's a local PID which is not a field. -/ -class DiscreteValuationRing (R : Type u) [CommRing R] [IsDomain R] +class IsDiscreteValuationRing (R : Type u) [CommRing R] [IsDomain R] extends IsPrincipalIdealRing R, IsLocalRing R : Prop where not_a_field' : maximalIdeal R ≠ ⊥ -namespace DiscreteValuationRing +namespace IsDiscreteValuationRing -variable (R : Type u) [CommRing R] [IsDomain R] [DiscreteValuationRing R] +variable (R : Type u) [CommRing R] [IsDomain R] [IsDiscreteValuationRing R] theorem not_a_field : maximalIdeal R ≠ ⊥ := not_a_field' @@ -104,7 +104,7 @@ theorem exists_prime : ∃ ϖ : R, Prime ϖ := /-- An integral domain is a DVR iff it's a PID with a unique non-zero prime ideal. -/ theorem iff_pid_with_one_nonzero_prime (R : Type u) [CommRing R] [IsDomain R] : - DiscreteValuationRing R ↔ IsPrincipalIdealRing R ∧ ∃! P : Ideal R, P ≠ ⊥ ∧ IsPrime P := by + IsDiscreteValuationRing R ↔ IsPrincipalIdealRing R ∧ ∃! P : Ideal R, P ≠ ⊥ ∧ IsPrime P := by constructor · intro RDVR rcases id RDVR with ⟨Rlocal⟩ @@ -137,9 +137,9 @@ theorem associated_of_irreducible {a b : R} (ha : Irreducible a) (hb : Irreducib rw [irreducible_iff_uniformizer] at ha hb rw [← span_singleton_eq_span_singleton, ← ha, hb] -end DiscreteValuationRing +end IsDiscreteValuationRing -namespace DiscreteValuationRing +namespace IsDiscreteValuationRing variable (R : Type*) @@ -179,7 +179,7 @@ variable [IsDomain R] /-- An integral domain in which there is an irreducible element `p` such that every nonzero element is associated to a power of `p` is a unique factorization domain. -See `DiscreteValuationRing.ofHasUnitMulPowIrreducibleFactorization`. -/ +See `IsDiscreteValuationRing.ofHasUnitMulPowIrreducibleFactorization`. -/ theorem toUniqueFactorizationMonoid (hR : HasUnitMulPowIrreducibleFactorization R) : UniqueFactorizationMonoid R := let p := Classical.choose hR @@ -269,7 +269,7 @@ is a discrete valuation ring. theorem of_ufd_of_unique_irreducible {R : Type u} [CommRing R] [IsDomain R] [UniqueFactorizationMonoid R] (h₁ : ∃ p : R, Irreducible p) (h₂ : ∀ ⦃p q : R⦄, Irreducible p → Irreducible q → Associated p q) : - DiscreteValuationRing R := by + IsDiscreteValuationRing R := by rw [iff_pid_with_one_nonzero_prime] haveI PID : IsPrincipalIdealRing R := aux_pid_of_ufd_of_unique_irreducible R h₁ h₂ obtain ⟨p, hp⟩ := h₁ @@ -290,7 +290,7 @@ such that every nonzero element is associated to a power of `p` is a discrete valuation ring. -/ theorem ofHasUnitMulPowIrreducibleFactorization {R : Type u} [CommRing R] [IsDomain R] - (hR : HasUnitMulPowIrreducibleFactorization R) : DiscreteValuationRing R := by + (hR : HasUnitMulPowIrreducibleFactorization R) : IsDiscreteValuationRing R := by letI : UniqueFactorizationMonoid R := hR.toUniqueFactorizationMonoid apply of_ufd_of_unique_irreducible _ hR.unique_irreducible obtain ⟨p, hp, H⟩ := hR @@ -298,7 +298,7 @@ theorem ofHasUnitMulPowIrreducibleFactorization {R : Type u} [CommRing R] [IsDom section -variable [CommRing R] [IsDomain R] [DiscreteValuationRing R] +variable [CommRing R] [IsDomain R] [IsDiscreteValuationRing R] variable {R} theorem associated_pow_irreducible {x : R} (hx : x ≠ 0) {ϖ : R} (hirr : Irreducible ϖ) : @@ -369,7 +369,7 @@ theorem unit_mul_pow_congr_unit {ϖ : R} (hirr : Irreducible ϖ) (u v : Rˣ) (m open Classical in /-- The `ℕ∞`-valued additive valuation on a DVR. -/ -noncomputable def addVal (R : Type u) [CommRing R] [IsDomain R] [DiscreteValuationRing R] : +noncomputable def addVal (R : Type u) [CommRing R] [IsDomain R] [IsDiscreteValuationRing R] : AddValuation R ℕ∞ := multiplicity_addValuation (Classical.choose_spec (exists_prime R)) @@ -381,6 +381,7 @@ theorem addVal_def (r : R) (u : Rˣ) {ϖ : R} (hϖ : Irreducible ϖ) (n : ℕ) ( emultiplicity_eq_of_associated_right (Associated.symm ⟨u, mul_comm _ _⟩), emultiplicity_pow_self_of_prime (irreducible_iff_prime.1 hϖ)] +/-- An alternative definition of the additive valuation, taking units into account.-/ theorem addVal_def' (u : Rˣ) {ϖ : R} (hϖ : Irreducible ϖ) (n : ℕ) : addVal R ((u : R) * ϖ ^ n) = n := addVal_def _ u hϖ n rfl @@ -439,7 +440,7 @@ theorem addVal_add {a b : R} : min (addVal R a) (addVal R b) ≤ addVal R (a + b end -instance (R : Type*) [CommRing R] [IsDomain R] [DiscreteValuationRing R] : +instance (R : Type*) [CommRing R] [IsDomain R] [IsDiscreteValuationRing R] : IsHausdorff (maximalIdeal R) R where haus' x hx := by obtain ⟨ϖ, hϖ⟩ := exists_irreducible R @@ -447,14 +448,14 @@ instance (R : Type*) [CommRing R] [IsDomain R] [DiscreteValuationRing R] : Ideal.span_singleton_pow, Ideal.mem_span_singleton, ← addVal_le_iff_dvd, hϖ.addVal_pow] at hx rwa [← addVal_eq_top_iff, ← WithTop.forall_ge_iff_eq_top] -end DiscreteValuationRing +end IsDiscreteValuationRing section -variable (A : Type u) [CommRing A] [IsDomain A] [DiscreteValuationRing A] +variable (A : Type u) [CommRing A] [IsDomain A] [IsDiscreteValuationRing A] /-- A DVR is a valuation ring. -/ -instance (priority := 100) of_discreteValuationRing : ValuationRing A := inferInstance +instance (priority := 100) of_isDiscreteValuationRing : ValuationRing A := inferInstance end diff --git a/Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean b/Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean index b00298b9d3929..26906e48f8d9c 100644 --- a/Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean +++ b/Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean @@ -12,7 +12,7 @@ import Mathlib.RingTheory.Ideal.Cotangent # Equivalent conditions for DVR -In `DiscreteValuationRing.TFAE`, we show that the following are equivalent for a +In `IsDiscreteValuationRing.TFAE`, we show that the following are equivalent for a noetherian local domain that is not a field `(R, m, k)`: - `R` is a discrete valuation ring - `R` is a valuation ring @@ -47,7 +47,7 @@ theorem exists_maximalIdeal_pow_eq_of_principal [IsNoetherianRing R] [IsLocalRin rintro rfl apply Ring.ne_bot_of_isMaximal_of_not_isField (maximalIdeal.isMaximal R) h simp [hx] - have hx' := DiscreteValuationRing.irreducible_of_span_eq_maximalIdeal x this hx + have hx' := IsDiscreteValuationRing.irreducible_of_span_eq_maximalIdeal x this hx have H' : ∀ r : R, r ≠ 0 → r ∈ nonunits R → ∃ n : ℕ, Associated (x ^ n) r := by intro r hr₁ hr₂ obtain ⟨f, hf₁, rfl, hf₂⟩ := (WfDvdMonoid.not_unit_iff_exists_factors_eq r hr₁).mp hr₂ @@ -157,7 +157,7 @@ The following are equivalent: 5. `dimₖ m/m² ≤ 1` 6. Every nonzero ideal is a power of `m`. -Also see `DiscreteValuationRing.TFAE` for a version assuming `¬ IsField R`. +Also see `IsDiscreteValuationRing.TFAE` for a version assuming `¬ IsField R`. -/ theorem tfae_of_isNoetherianRing_of_isLocalRing_of_isDomain [IsNoetherianRing R] [IsLocalRing R] [IsDomain R] : @@ -207,10 +207,10 @@ noetherian local domain that is not a field `(R, m, k)`: Also see `tfae_of_isNoetherianRing_of_isLocalRing_of_isDomain` for a version without `¬ IsField R`. -/ -theorem DiscreteValuationRing.TFAE [IsNoetherianRing R] [IsLocalRing R] [IsDomain R] +theorem IsDiscreteValuationRing.TFAE [IsNoetherianRing R] [IsLocalRing R] [IsDomain R] (h : ¬IsField R) : List.TFAE - [DiscreteValuationRing R, ValuationRing R, IsDedekindDomain R, + [IsDiscreteValuationRing R, ValuationRing R, IsDedekindDomain R, IsIntegrallyClosed R ∧ ∃! P : Ideal R, P ≠ ⊥ ∧ P.IsPrime, (maximalIdeal R).IsPrincipal, finrank (ResidueField R) (CotangentSpace R) = 1, ∀ (I) (_ : I ≠ ⊥), ∃ n : ℕ, I = maximalIdeal R ^ n] := by @@ -227,19 +227,19 @@ theorem DiscreteValuationRing.TFAE [IsNoetherianRing R] [IsLocalRing R] [IsDomai variable {R} lemma IsLocalRing.finrank_CotangentSpace_eq_one_iff [IsNoetherianRing R] [IsLocalRing R] - [IsDomain R] : finrank (ResidueField R) (CotangentSpace R) = 1 ↔ DiscreteValuationRing R := by + [IsDomain R] : finrank (ResidueField R) (CotangentSpace R) = 1 ↔ IsDiscreteValuationRing R := by by_cases hR : IsField R · letI := hR.toField simp only [finrank_cotangentSpace_eq_zero, zero_ne_one, false_iff] exact fun h ↦ h.3 maximalIdeal_eq_bot - · exact (DiscreteValuationRing.TFAE R hR).out 5 0 + · exact (IsDiscreteValuationRing.TFAE R hR).out 5 0 @[deprecated (since := "2024-11-09")] alias LocalRing.finrank_CotangentSpace_eq_one_iff := IsLocalRing.finrank_CotangentSpace_eq_one_iff variable (R) -lemma IsLocalRing.finrank_CotangentSpace_eq_one [IsDomain R] [DiscreteValuationRing R] : +lemma IsLocalRing.finrank_CotangentSpace_eq_one [IsDomain R] [IsDiscreteValuationRing R] : finrank (ResidueField R) (CotangentSpace R) = 1 := finrank_CotangentSpace_eq_one_iff.mpr ‹_› diff --git a/Mathlib/RingTheory/PowerSeries/Inverse.lean b/Mathlib/RingTheory/PowerSeries/Inverse.lean index 419143bde5cd0..8a1b967d6eeca 100644 --- a/Mathlib/RingTheory/PowerSeries/Inverse.lean +++ b/Mathlib/RingTheory/PowerSeries/Inverse.lean @@ -284,11 +284,11 @@ instance : IsLocalRing R⟦X⟧ := end IsLocalRing -section DiscreteValuationRing +section IsDiscreteValuationRing variable {k : Type*} [Field k] -open DiscreteValuationRing +open IsDiscreteValuationRing theorem hasUnitMulPowIrreducibleFactorization : HasUnitMulPowIrreducibleFactorization k⟦X⟧ := @@ -303,7 +303,7 @@ theorem hasUnitMulPowIrreducibleFactorization : instance : UniqueFactorizationMonoid k⟦X⟧ := hasUnitMulPowIrreducibleFactorization.toUniqueFactorizationMonoid -instance : DiscreteValuationRing k⟦X⟧ := +instance : IsDiscreteValuationRing k⟦X⟧ := ofHasUnitMulPowIrreducibleFactorization hasUnitMulPowIrreducibleFactorization instance isNoetherianRing : IsNoetherianRing k⟦X⟧ := @@ -376,7 +376,7 @@ def residueFieldOfPowerSeries : ResidueField k⟦X⟧ ≃+* k := (Ideal.quotEquivOfEq (ker_coeff_eq_max_ideal).symm).trans (RingHom.quotientKerEquivOfSurjective constantCoeff_surj) -end DiscreteValuationRing +end IsDiscreteValuationRing end PowerSeries diff --git a/Mathlib/RingTheory/Valuation/ValuationRing.lean b/Mathlib/RingTheory/Valuation/ValuationRing.lean index aa068f34c3408..d0ff48e43af4d 100644 --- a/Mathlib/RingTheory/Valuation/ValuationRing.lean +++ b/Mathlib/RingTheory/Valuation/ValuationRing.lean @@ -41,7 +41,7 @@ The `ValuationRing` class is kept to be in sync with the literature. -/ -assert_not_exists DiscreteValuationRing +assert_not_exists IsDiscreteValuationRing universe u v w diff --git a/Mathlib/RingTheory/WittVector/DiscreteValuationRing.lean b/Mathlib/RingTheory/WittVector/DiscreteValuationRing.lean index 7261a54d9e359..459ea21faac7a 100644 --- a/Mathlib/RingTheory/WittVector/DiscreteValuationRing.lean +++ b/Mathlib/RingTheory/WittVector/DiscreteValuationRing.lean @@ -21,7 +21,7 @@ When `k` is also a field, this `b` can be chosen to be a unit of `𝕎 k`. * `WittVector.exists_eq_pow_p_mul`: the existence of this element `b` over a perfect ring * `WittVector.exists_eq_pow_p_mul'`: the existence of this unit `b` over a perfect field -* `WittVector.discreteValuationRing`: `𝕎 k` is a discrete valuation ring if `k` is a perfect field +* `WittVector.isDiscreteValuationRing`: `𝕎 k` is a discrete valuation ring if `k` is a perfect field -/ @@ -147,8 +147,8 @@ https://github.com/leanprover/lean4/issues/1102 -/ /-- The ring of Witt Vectors of a perfect field of positive characteristic is a DVR. -/ -theorem discreteValuationRing : DiscreteValuationRing (𝕎 k) := - DiscreteValuationRing.ofHasUnitMulPowIrreducibleFactorization (by +theorem isDiscreteValuationRing : IsDiscreteValuationRing (𝕎 k) := + IsDiscreteValuationRing.ofHasUnitMulPowIrreducibleFactorization (by refine ⟨p, irreducible p, fun {x} hx => ?_⟩ obtain ⟨n, b, hb⟩ := exists_eq_pow_p_mul' x hx exact ⟨n, b, hb.symm⟩) From bb0575e329e059f7ec31c43337a2094282f6be3c Mon Sep 17 00:00:00 2001 From: Lorenzo Luccioli Date: Tue, 17 Dec 2024 13:46:44 +0000 Subject: [PATCH 755/829] feat(Data/EReal): Add results about `EReal` (#17087) Add results about operations on the extended real numbers. - Instances for `EReal`: `CharZero` and `NoZeroDivisors`. - Results about the coercion to the reals: `toReal_eq_zero_iff`, `toReal_ne_zero_iff`, `toReal_nonneg`, `toReal_nonpos`, `coe_ennreal_toReal`. - Results about the addition of EReals: `add_ne_bot_iff`, `bot_lt_add_iff`, `add_ne_top` and similar. - Reorganize and rename the lemmas `le_neg` and similar to make them more consistent, add the missing ones. - Results about the subtraction of EReals: `top_sub_of_ne_top`, `sub_self`, `sub_self_le_zero`, `sub_nonneg` and similar, `sub_add_cancel` and similar. - Results about the multiplication of Ereals: `top_mul_coe_ennreal`, `coe_ennreal_mul_top`, `mul_pos_iff` and similar, `mul_eq_top` and similar, `coe_mul_add_of_nonneg` and other distributivity lemmas, `nsmul_eq_mul`. - Move and simplify proof of `mul_pos`. - Clenaup. --- Mathlib/Data/Real/EReal.lean | 305 +++++++++++++++++++++----- Mathlib/Topology/Instances/EReal.lean | 4 +- 2 files changed, 255 insertions(+), 54 deletions(-) diff --git a/Mathlib/Data/Real/EReal.lean b/Mathlib/Data/Real/EReal.lean index 723fbe073737f..c1f2ec68193c6 100644 --- a/Mathlib/Data/Real/EReal.lean +++ b/Mathlib/Data/Real/EReal.lean @@ -13,8 +13,8 @@ import Mathlib.Data.Sign This file defines `EReal`, the real numbers together with a top and bottom element, referred to as ⊤ and ⊥. It is implemented as `WithBot (WithTop ℝ)` -Addition and multiplication are problematic in the presence of ±∞, but -negation has a natural definition and satisfies the usual properties. +Addition and multiplication are problematic in the presence of ±∞, but negation has a natural +definition and satisfies the usual properties, in particular it is an order reversing isomorphism An ad hoc addition is defined, for which `EReal` is an `AddCommMonoid`, and even an ordered one (if `a ≤ a'` and `b ≤ b'` then `a + b ≤ a' + b'`). @@ -29,6 +29,9 @@ but it is sometimes convenient to have. An ad hoc multiplication is defined, for which `EReal` is a `CommMonoidWithZero`. We make the choice that `0 * x = x * 0 = 0` for any `x` (while the other cases are defined non-ambiguously). This does not distribute with addition, as `⊥ = ⊥ + ⊤ = 1*⊥ + (-1)*⊥ ≠ (1 - 1) * ⊥ = 0 * ⊥ = 0`. +Distributivity `x * (y + z) = x * y + x * z` is recovered in the case where either `0 ≤ x < ⊤`, +see `Ereal.left_distrib_of_nonneg_of_ne_top`, or `0 ≤ y, z`, see `Ereal.left_distrib_of_nonneg` +(similarily for right distributivity). `EReal` is a `CompleteLinearOrder`; this is deduced by type class inference from the fact that `WithBot (WithTop L)` is a complete linear order if `L` is @@ -71,6 +74,8 @@ instance : AddCommMonoidWithOne EReal := instance : DenselyOrdered EReal := inferInstanceAs (DenselyOrdered (WithBot (WithTop ℝ))) +instance : CharZero EReal := inferInstanceAs (CharZero (WithBot (WithTop ℝ))) + /-- The canonical inclusion from reals to ereals. Registered as a coercion. -/ @[coe] def Real.toEReal : ℝ → EReal := some ∘ some @@ -106,7 +111,7 @@ protected theorem coe_eq_coe_iff {x y : ℝ} : (x : EReal) = (y : EReal) ↔ x = protected theorem coe_ne_coe_iff {x y : ℝ} : (x : EReal) ≠ (y : EReal) ↔ x ≠ y := coe_injective.ne_iff -/-- The canonical map from nonnegative extended reals to extended reals -/ +/-- The canonical map from nonnegative extended reals to extended reals. -/ @[coe] def _root_.ENNReal.toEReal : ℝ≥0∞ → EReal | ⊤ => ⊤ | .some x => x.1 @@ -357,6 +362,31 @@ protected theorem coe_pos {x : ℝ} : (0 : EReal) < x ↔ 0 < x := protected theorem coe_neg' {x : ℝ} : (x : EReal) < 0 ↔ x < 0 := EReal.coe_lt_coe_iff +lemma toReal_eq_zero_iff {x : EReal} : x.toReal = 0 ↔ x = 0 ∨ x = ⊤ ∨ x = ⊥ := by + cases x <;> norm_num + +lemma toReal_ne_zero_iff {x : EReal} : x.toReal ≠ 0 ↔ x ≠ 0 ∧ x ≠ ⊤ ∧ x ≠ ⊥ := by + simp only [ne_eq, toReal_eq_zero_iff, not_or] + +lemma toReal_eq_toReal {x y : EReal} (hx_top : x ≠ ⊤) (hx_bot : x ≠ ⊥) + (hy_top : y ≠ ⊤) (hy_bot : y ≠ ⊥) : + x.toReal = y.toReal ↔ x = y := by + lift x to ℝ using ⟨hx_top, hx_bot⟩ + lift y to ℝ using ⟨hy_top, hy_bot⟩ + simp + +lemma toReal_nonneg {x : EReal} (hx : 0 ≤ x) : 0 ≤ x.toReal := by + cases x + · norm_num + · exact toReal_coe _ ▸ EReal.coe_nonneg.mp hx + · norm_num + +lemma toReal_nonpos {x : EReal} (hx : x ≤ 0) : x.toReal ≤ 0 := by + cases x + · norm_num + · exact toReal_coe _ ▸ EReal.coe_nonpos.mp hx + · norm_num + theorem toReal_le_toReal {x y : EReal} (h : x ≤ y) (hx : x ≠ ⊥) (hy : y ≠ ⊤) : x.toReal ≤ y.toReal := by lift x to ℝ using ⟨ne_top_of_le_ne_top hy h, hx⟩ @@ -542,6 +572,10 @@ theorem toReal_coe_ennreal : ∀ {x : ℝ≥0∞}, toReal (x : EReal) = ENNReal. theorem coe_ennreal_ofReal {x : ℝ} : (ENNReal.ofReal x : EReal) = max x 0 := rfl +lemma coe_ennreal_toReal {x : ℝ≥0∞} (hx : x ≠ ∞) : (x.toReal : EReal) = x := by + lift x to ℝ≥0 using hx + rfl + theorem coe_nnreal_eq_coe_real (x : ℝ≥0) : ((x : ℝ≥0∞) : EReal) = (x : ℝ) := rfl @@ -657,18 +691,18 @@ theorem natCast_ne_bot (n : ℕ) : (n : EReal) ≠ ⊥ := Ne.symm (ne_of_beq_fal theorem natCast_ne_top (n : ℕ) : (n : EReal) ≠ ⊤ := Ne.symm (ne_of_beq_false rfl) -@[simp, norm_cast] +@[norm_cast] theorem natCast_eq_iff {m n : ℕ} : (m : EReal) = (n : EReal) ↔ m = n := by rw [← coe_coe_eq_natCast n, ← coe_coe_eq_natCast m, EReal.coe_eq_coe_iff, Nat.cast_inj] theorem natCast_ne_iff {m n : ℕ} : (m : EReal) ≠ (n : EReal) ↔ m ≠ n := not_iff_not.2 natCast_eq_iff -@[simp, norm_cast] +@[norm_cast] theorem natCast_le_iff {m n : ℕ} : (m : EReal) ≤ (n : EReal) ↔ m ≤ n := by rw [← coe_coe_eq_natCast n, ← coe_coe_eq_natCast m, EReal.coe_le_coe_iff, Nat.cast_le] -@[simp, norm_cast] +@[norm_cast] theorem natCast_lt_iff {m n : ℕ} : (m : EReal) < (n : EReal) ↔ m < n := by rw [← coe_coe_eq_natCast n, ← coe_coe_eq_natCast m, EReal.coe_lt_coe_iff, Nat.cast_lt] @@ -727,9 +761,11 @@ theorem bot_add (x : EReal) : ⊥ + x = ⊥ := theorem add_eq_bot_iff {x y : EReal} : x + y = ⊥ ↔ x = ⊥ ∨ y = ⊥ := WithBot.add_eq_bot +lemma add_ne_bot_iff {x y : EReal} : x + y ≠ ⊥ ↔ x ≠ ⊥ ∧ y ≠ ⊥ := WithBot.add_ne_bot + @[simp] theorem bot_lt_add_iff {x y : EReal} : ⊥ < x + y ↔ ⊥ < x ∧ ⊥ < y := by - simp [bot_lt_iff_ne_bot, not_or] + simp [bot_lt_iff_ne_bot] @[simp] theorem top_add_top : (⊤ : EReal) + ⊤ = ⊤ := @@ -828,9 +864,26 @@ theorem add_lt_add_of_lt_of_le {x y z t : EReal} (h : x < y) (h' : z ≤ t) (hz (ht : t ≠ ⊤) : x + z < y + t := add_lt_add_of_lt_of_le' h h' (ne_bot_of_le_ne_bot hz h') fun ht' => (ht ht').elim -theorem add_lt_top {x y : EReal} (hx : x ≠ ⊤) (hy : y ≠ ⊤) : x + y < ⊤ := by - rw [← EReal.top_add_top] - exact EReal.add_lt_add hx.lt_top hy.lt_top +theorem add_lt_top {x y : EReal} (hx : x ≠ ⊤) (hy : y ≠ ⊤) : x + y < ⊤ := + add_lt_add hx.lt_top hy.lt_top + +lemma add_ne_top {x y : EReal} (hx : x ≠ ⊤) (hy : y ≠ ⊤) : x + y ≠ ⊤ := + lt_top_iff_ne_top.mp <| add_lt_top hx hy + +lemma add_ne_top_iff_ne_top₂ {x y : EReal} (hx : x ≠ ⊥) (hy : y ≠ ⊥) : + x + y ≠ ⊤ ↔ x ≠ ⊤ ∧ y ≠ ⊤ := by + refine ⟨?_, fun h ↦ add_ne_top h.1 h.2⟩ + cases x <;> simp_all only [ne_eq, not_false_eq_true, top_add_of_ne_bot, not_true_eq_false, + IsEmpty.forall_iff] + cases y <;> simp_all only [not_false_eq_true, ne_eq, add_top_of_ne_bot, not_true_eq_false, + coe_ne_top, and_self, implies_true] + +lemma add_ne_top_iff_ne_top_left {x y : EReal} (hy : y ≠ ⊥) (hy' : y ≠ ⊤) : + x + y ≠ ⊤ ↔ x ≠ ⊤ := by + cases x <;> simp [add_ne_top_iff_ne_top₂, hy, hy'] + +lemma add_ne_top_iff_ne_top_right {x y : EReal} (hx : x ≠ ⊥) (hx' : x ≠ ⊤) : + x + y ≠ ⊤ ↔ y ≠ ⊤ := add_comm x y ▸ add_ne_top_iff_ne_top_left hx hx' /-- We do not have a notion of `LinearOrderedAddCommMonoidWithBot` but we can at least make the order dual of the extended reals into a `LinearOrderedAddCommMonoidWithTop`. -/ @@ -905,34 +958,41 @@ theorem neg_strictAnti : StrictAnti (- · : EReal → EReal) := @[simp] theorem neg_lt_neg_iff {a b : EReal} : -a < -b ↔ b < a := neg_strictAnti.lt_iff_lt -/-- `-a ≤ b ↔ -b ≤ a` on `EReal`. -/ +/-- `-a ≤ b` if and only if `-b ≤ a` on `EReal`. -/ protected theorem neg_le {a b : EReal} : -a ≤ b ↔ -b ≤ a := by rw [← neg_le_neg_iff, neg_neg] -/-- if `-a ≤ b` then `-b ≤ a` on `EReal`. -/ +/-- If `-a ≤ b` then `-b ≤ a` on `EReal`. -/ protected theorem neg_le_of_neg_le {a b : EReal} (h : -a ≤ b) : -b ≤ a := EReal.neg_le.mp h -/-- `a ≤ -b → b ≤ -a` on ereal -/ -theorem le_neg_of_le_neg {a b : EReal} (h : a ≤ -b) : b ≤ -a := by - rwa [← neg_neg b, EReal.neg_le, neg_neg] +/-- `a ≤ -b` if and only if `b ≤ -a` on `EReal`. -/ +protected theorem le_neg {a b : EReal} : a ≤ -b ↔ b ≤ -a := by + rw [← neg_le_neg_iff, neg_neg] -/-- Negation as an order reversing isomorphism on `EReal`. -/ -def negOrderIso : EReal ≃o ERealᵒᵈ := - { Equiv.neg EReal with - toFun := fun x => OrderDual.toDual (-x) - invFun := fun x => -OrderDual.ofDual x - map_rel_iff' := neg_le_neg_iff } +/-- If `a ≤ -b` then `b ≤ -a` on `EReal`. -/ +protected theorem le_neg_of_le_neg {a b : EReal} (h : a ≤ -b) : b ≤ -a := EReal.le_neg.mp h +/-- `-a < b` if and only if `-b < a` on `EReal`. -/ theorem neg_lt_comm {a b : EReal} : -a < b ↔ -b < a := by rw [← neg_lt_neg_iff, neg_neg] @[deprecated (since := "2024-11-19")] alias neg_lt_iff_neg_lt := neg_lt_comm -theorem neg_lt_of_neg_lt {a b : EReal} (h : -a < b) : -b < a := neg_lt_comm.1 h +/-- If `-a < b` then `-b < a` on `EReal`. -/ +protected theorem neg_lt_of_neg_lt {a b : EReal} (h : -a < b) : -b < a := neg_lt_comm.mp h +/-- `-a < b` if and only if `-b < a` on `EReal`. -/ theorem lt_neg_comm {a b : EReal} : a < -b ↔ b < -a := by rw [← neg_lt_neg_iff, neg_neg] -theorem lt_neg_of_lt_neg {a b : EReal} (h : a < -b) : b < -a := lt_neg_comm.1 h +/-- If `a < -b` then `b < -a` on `EReal`. -/ +protected theorem lt_neg_of_lt_neg {a b : EReal} (h : a < -b) : b < -a := lt_neg_comm.mp h + +/-- Negation as an order reversing isomorphism on `EReal`. -/ +def negOrderIso : EReal ≃o ERealᵒᵈ := + { Equiv.neg EReal with + toFun := fun x => OrderDual.toDual (-x) + invFun := fun x => -OrderDual.ofDual x + map_rel_iff' := neg_le_neg_iff } lemma neg_add {x y : EReal} (h1 : x ≠ ⊥ ∨ y ≠ ⊤) (h2 : x ≠ ⊤ ∨ y ≠ ⊥) : - (x + y) = - x - y := by @@ -971,11 +1031,34 @@ theorem top_sub_coe (x : ℝ) : (⊤ : EReal) - x = ⊤ := theorem coe_sub_bot (x : ℝ) : (x : EReal) - ⊥ = ⊤ := rfl -lemma sub_bot {a : EReal} (h : a ≠ ⊥) : a - ⊥ = ⊤ := by - induction a - · simp only [ne_eq, not_true_eq_false] at h - · rw [coe_sub_bot] - · rw [top_sub_bot] +@[simp] +lemma sub_bot {x : EReal} (h : x ≠ ⊥) : x - ⊥ = ⊤ := by + cases x <;> tauto + +@[simp] +lemma top_sub {x : EReal} (hx : x ≠ ⊤) : ⊤ - x = ⊤ := by + cases x <;> tauto + +@[simp] +lemma sub_self {x : EReal} (h_top : x ≠ ⊤) (h_bot : x ≠ ⊥) : x - x = 0 := by + cases x <;> simp_all [← coe_sub] + +lemma sub_self_le_zero {x : EReal} : x - x ≤ 0 := by + cases x <;> simp + +lemma sub_nonneg {x y : EReal} (h_top : x ≠ ⊤ ∨ y ≠ ⊤) (h_bot : x ≠ ⊥ ∨ y ≠ ⊥) : + 0 ≤ x - y ↔ y ≤ x := by + cases x <;> cases y <;> simp_all [← EReal.coe_sub] + +lemma sub_nonpos {x y : EReal} : x - y ≤ 0 ↔ x ≤ y := by + cases x <;> cases y <;> simp [← EReal.coe_sub] + +lemma sub_pos {x y : EReal} : 0 < x - y ↔ y < x := by + cases x <;> cases y <;> simp [← EReal.coe_sub] + +lemma sub_neg {x y : EReal} (h_top : x ≠ ⊤ ∨ y ≠ ⊤) (h_bot : x ≠ ⊥ ∨ y ≠ ⊥) : + x - y < 0 ↔ x < y := by + cases x <;> cases y <;> simp_all [← EReal.coe_sub] theorem sub_le_sub {x y z t : EReal} (h : x ≤ y) (h' : t ≤ z) : x - z ≤ y - t := add_le_add h (neg_le_neg_iff.2 h') @@ -1001,15 +1084,22 @@ theorem toReal_sub {x y : EReal} (hx : x ≠ ⊤) (h'x : x ≠ ⊥) (hy : y ≠ lift y to ℝ using ⟨hy, h'y⟩ rfl -lemma sub_add_cancel_left {a : EReal} {b : Real} : a - b + b = a := by - induction a - · rw [bot_sub b, bot_add b] - · norm_cast; linarith - · rw [top_sub_coe b, top_add_coe b] - lemma add_sub_cancel_right {a : EReal} {b : Real} : a + b - b = a := by - rw [sub_eq_add_neg, add_assoc, add_comm (b : EReal), ← add_assoc, ← sub_eq_add_neg] - exact sub_add_cancel_left + cases a <;> norm_cast + exact _root_.add_sub_cancel_right _ _ + +lemma add_sub_cancel_left {a : EReal} {b : Real} : b + a - b = a := by + rw [add_comm, EReal.add_sub_cancel_right] + +lemma sub_add_cancel {a : EReal} {b : Real} : a - b + b = a := by + rw [add_comm, ← add_sub_assoc, add_sub_cancel_left] + +lemma sub_add_cancel_right {a : EReal} {b : Real} : b - (a + b) = -a := by + cases a <;> norm_cast + exact _root_.sub_add_cancel_right _ _ + +lemma sub_add_cancel_left {a : EReal} {b : Real} : b - (b + a) = -a := by + rw [add_comm, sub_add_cancel_right] lemma le_sub_iff_add_le {a b c : EReal} (hb : b ≠ ⊥ ∨ c ≠ ⊥) (ht : b ≠ ⊤ ∨ c ≠ ⊤) : a ≤ c - b ↔ a + b ≤ c := by @@ -1018,7 +1108,7 @@ lemma le_sub_iff_add_le {a b c : EReal} (hb : b ≠ ⊥ ∨ c ≠ ⊥) (ht : b simp only [ne_eq, not_true_eq_false, false_or] at hb simp only [sub_bot hb, le_top, add_bot, bot_le] | h_real b => - rw [← (addLECancellable_coe b).add_le_add_iff_right, sub_add_cancel_left] + rw [← (addLECancellable_coe b).add_le_add_iff_right, sub_add_cancel] | h_top => simp only [ne_eq, not_true_eq_false, false_or, sub_top, le_bot_iff] at ht ⊢ refine ⟨fun h ↦ h ▸ (bot_add ⊤).symm ▸ bot_le, fun h ↦ ?_⟩ @@ -1093,9 +1183,9 @@ lemma add_le_of_forall_lt {a b c : EReal} (h : ∀ a' < a, ∀ b' < b, a' + b' lemma le_add_of_forall_gt {a b c : EReal} (h₁ : a ≠ ⊥ ∨ b ≠ ⊤) (h₂ : a ≠ ⊤ ∨ b ≠ ⊥) (h : ∀ a' > a, ∀ b' > b, c ≤ a' + b') : c ≤ a + b := by rw [← neg_le_neg_iff, neg_add h₁ h₂] - exact add_le_of_forall_lt fun a' ha' b' hb' ↦ le_neg_of_le_neg - <| (h (-a') (lt_neg_of_lt_neg ha') (-b') (lt_neg_of_lt_neg hb')).trans_eq - (neg_add (.inr hb'.ne_top) (.inl ha'.ne_top)).symm + refine add_le_of_forall_lt fun a' ha' b' hb' ↦ EReal.le_neg_of_le_neg ?_ + rw [neg_add (.inr hb'.ne_top) (.inl ha'.ne_top)] + exact h _ (EReal.lt_neg_of_lt_neg ha') _ (EReal.lt_neg_of_lt_neg hb') @[deprecated (since := "2024-11-19")] alias top_add_le_of_forall_add_le := add_le_of_forall_lt @[deprecated (since := "2024-11-19")] alias add_le_of_forall_add_le := add_le_of_forall_lt @@ -1145,20 +1235,16 @@ lemma top_mul_of_pos {x : EReal} (h : 0 < x) : ⊤ * x = ⊤ := by rw [EReal.mul_comm] exact mul_top_of_pos h -/-- The product of two positive extended real numbers is positive. -/ -lemma mul_pos {a b : EReal} (ha : 0 < a) (hb : 0 < b) : 0 < a * b := by - induction a - · exfalso; exact not_lt_bot ha - · induction b - · exfalso; exact not_lt_bot hb - · norm_cast at *; exact Left.mul_pos ha hb - · rw [EReal.mul_comm, top_mul_of_pos ha]; exact hb - · rw [top_mul_of_pos hb]; exact ha - lemma top_mul_of_neg {x : EReal} (h : x < 0) : ⊤ * x = ⊥ := by rw [EReal.mul_comm] exact mul_top_of_neg h +lemma top_mul_coe_ennreal {x : ℝ≥0∞} (hx : x ≠ 0) : ⊤ * (x : EReal) = ⊤ := + top_mul_of_pos <| coe_ennreal_pos.mpr <| pos_iff_ne_zero.mpr hx + +lemma coe_ennreal_mul_top {x : ℝ≥0∞} (hx : x ≠ 0) : (x : EReal) * ⊤ = ⊤ := by + rw [EReal.mul_comm, top_mul_coe_ennreal hx] + lemma coe_mul_bot_of_pos {x : ℝ} (h : 0 < x) : (x : EReal) * ⊥ = ⊥ := if_pos h @@ -1199,6 +1285,43 @@ lemma toReal_mul {x y : EReal} : toReal (x * y) = toReal x * toReal y := by | pos_bot _ h => simp [coe_mul_bot_of_pos h] | neg_bot _ h => simp [coe_mul_bot_of_neg h] +instance : NoZeroDivisors EReal where + eq_zero_or_eq_zero_of_mul_eq_zero := by + intro a b h + contrapose! h + cases a <;> cases b <;> try {· simp_all [← EReal.coe_mul]} + · rcases lt_or_gt_of_ne h.2 with (h | h) + <;> simp [EReal.bot_mul_of_neg, EReal.bot_mul_of_pos, h] + · rcases lt_or_gt_of_ne h.1 with (h | h) + <;> simp [EReal.mul_bot_of_pos, EReal.mul_bot_of_neg, h] + · rcases lt_or_gt_of_ne h.1 with (h | h) + <;> simp [EReal.mul_top_of_neg, EReal.mul_top_of_pos, h] + · rcases lt_or_gt_of_ne h.2 with (h | h) + <;> simp [EReal.top_mul_of_pos, EReal.top_mul_of_neg, h] + +lemma mul_pos_iff {a b : EReal} : 0 < a * b ↔ 0 < a ∧ 0 < b ∨ a < 0 ∧ b < 0 := by + induction a, b using EReal.induction₂_symm with + | symm h => simp [EReal.mul_comm, h, and_comm] + | top_top => simp + | top_pos _ hx => simp [EReal.top_mul_coe_of_pos hx, hx] + | top_zero => simp + | top_neg _ hx => simp [hx, EReal.top_mul_coe_of_neg hx, le_of_lt] + | top_bot => simp + | pos_bot _ hx => simp [hx, EReal.coe_mul_bot_of_pos hx, le_of_lt] + | coe_coe x y => simp [← coe_mul, _root_.mul_pos_iff] + | zero_bot => simp + | neg_bot _ hx => simp [hx, EReal.coe_mul_bot_of_neg hx] + | bot_bot => simp + +lemma mul_nonneg_iff {a b : EReal} : 0 ≤ a * b ↔ 0 ≤ a ∧ 0 ≤ b ∨ a ≤ 0 ∧ b ≤ 0 := by + simp_rw [le_iff_lt_or_eq, mul_pos_iff, zero_eq_mul (a := a)] + rcases lt_trichotomy a 0 with (h | h | h) <;> rcases lt_trichotomy b 0 with (h' | h' | h') + <;> simp only [h, h', true_or, true_and, or_true, and_true] <;> tauto + +/-- The product of two positive extended real numbers is positive. -/ +lemma mul_pos {a b : EReal} (ha : 0 < a) (hb : 0 < b) : 0 < a * b := + mul_pos_iff.mpr (Or.inl ⟨ha, hb⟩) + /-- Induct on two ereals by performing case splits on the sign of one whenever the other is infinite. This version eliminates some cases by assuming that `P x y` implies `P (-x) y` for all `x`, `y`. -/ @@ -1251,6 +1374,60 @@ instance : HasDistribNeg EReal where rw [x.mul_comm, x.mul_comm] exact y.neg_mul x +lemma mul_neg_iff {a b : EReal} : a * b < 0 ↔ 0 < a ∧ b < 0 ∨ a < 0 ∧ 0 < b := by + nth_rw 1 [← neg_zero] + rw [lt_neg_comm, ← mul_neg a, mul_pos_iff, neg_lt_comm, lt_neg_comm, neg_zero] + +lemma mul_nonpos_iff {a b : EReal} : a * b ≤ 0 ↔ 0 ≤ a ∧ b ≤ 0 ∨ a ≤ 0 ∧ 0 ≤ b := by + nth_rw 1 [← neg_zero] + rw [EReal.le_neg, ← mul_neg, mul_nonneg_iff, EReal.neg_le, EReal.le_neg, neg_zero] + +lemma mul_eq_top (a b : EReal) : + a * b = ⊤ ↔ (a = ⊥ ∧ b < 0) ∨ (a < 0 ∧ b = ⊥) ∨ (a = ⊤ ∧ 0 < b) ∨ (0 < a ∧ b = ⊤) := by + induction a, b using EReal.induction₂_symm with + | symm h => + rw [EReal.mul_comm, h] + refine ⟨fun H ↦ ?_, fun H ↦ ?_⟩ <;> + cases H with + | inl h => exact Or.inr (Or.inl ⟨h.2, h.1⟩) + | inr h => cases h with + | inl h => exact Or.inl ⟨h.2, h.1⟩ + | inr h => cases h with + | inl h => exact Or.inr (Or.inr (Or.inr ⟨h.2, h.1⟩)) + | inr h => exact Or.inr (Or.inr (Or.inl ⟨h.2, h.1⟩)) + | top_top => simp + | top_pos _ hx => simp [EReal.top_mul_coe_of_pos hx, hx] + | top_zero => simp + | top_neg _ hx => simp [hx.le, EReal.top_mul_coe_of_neg hx] + | top_bot => simp + | pos_bot _ hx => simp [hx.le, EReal.coe_mul_bot_of_pos hx] + | coe_coe x y => + simpa only [EReal.coe_ne_bot, EReal.coe_neg', false_and, and_false, EReal.coe_ne_top, + EReal.coe_pos, or_self, iff_false, EReal.coe_mul] using EReal.coe_ne_top _ + | zero_bot => simp + | neg_bot _ hx => simp [hx, EReal.coe_mul_bot_of_neg hx] + | bot_bot => simp + +lemma mul_ne_top (a b : EReal) : + a * b ≠ ⊤ ↔ (a ≠ ⊥ ∨ 0 ≤ b) ∧ (0 ≤ a ∨ b ≠ ⊥) ∧ (a ≠ ⊤ ∨ b ≤ 0) ∧ (a ≤ 0 ∨ b ≠ ⊤) := by + rw [ne_eq, mul_eq_top] + -- push the negation while keeping the disjunctions, that is converting `¬(p ∧ q)` into `¬p ∨ ¬q` + -- rather than `p → ¬q`, since we already have disjunctions in the rhs + set_option push_neg.use_distrib true in push_neg + rfl + +lemma mul_eq_bot (a b : EReal) : + a * b = ⊥ ↔ (a = ⊥ ∧ 0 < b) ∨ (0 < a ∧ b = ⊥) ∨ (a = ⊤ ∧ b < 0) ∨ (a < 0 ∧ b = ⊤) := by + rw [← neg_eq_top_iff, ← EReal.neg_mul, mul_eq_top, neg_eq_bot_iff, neg_eq_top_iff, + neg_lt_comm, lt_neg_comm, neg_zero] + tauto + +lemma mul_ne_bot (a b : EReal) : + a * b ≠ ⊥ ↔ (a ≠ ⊥ ∨ b ≤ 0) ∧ (a ≤ 0 ∨ b ≠ ⊥) ∧ (a ≠ ⊤ ∨ 0 ≤ b) ∧ (0 ≤ a ∨ b ≠ ⊤) := by + rw [ne_eq, mul_eq_bot] + set_option push_neg.use_distrib true in push_neg + rfl + lemma right_distrib_of_nonneg {a b c : EReal} (ha : 0 ≤ a) (hb : 0 ≤ b) : (a + b) * c = a * c + b * c := by rcases eq_or_lt_of_le ha with (rfl | a_pos) @@ -1291,6 +1468,30 @@ lemma left_distrib_of_nonneg {a b c : EReal} (ha : 0 ≤ a) (hb : 0 ≤ b) : nth_rewrite 1 [EReal.mul_comm]; nth_rewrite 2 [EReal.mul_comm]; nth_rewrite 3 [EReal.mul_comm] exact right_distrib_of_nonneg ha hb +lemma left_distrib_of_nonneg_of_ne_top {x : EReal} (hx_nonneg : 0 ≤ x) + (hx_ne_top : x ≠ ⊤) (y z : EReal) : + x * (y + z) = x * y + x * z := by + by_cases hx0 : x = 0 + · simp [hx0] + replace hx0 : 0 < x := hx_nonneg.lt_of_ne' hx0 + lift x to ℝ using ⟨hx_ne_top, hx0.ne_bot⟩ + cases y <;> cases z <;> + simp [mul_bot_of_pos hx0, mul_top_of_pos hx0, ← coe_mul]; + rw_mod_cast [mul_add] + +lemma right_distrib_of_nonneg_of_ne_top {x : EReal} (hx_nonneg : 0 ≤ x) + (hx_ne_top : x ≠ ⊤) (y z : EReal) : + (y + z) * x = y * x + z * x := by + simpa only [EReal.mul_comm] using left_distrib_of_nonneg_of_ne_top hx_nonneg hx_ne_top y z + +@[simp] +lemma nsmul_eq_mul (n : ℕ) (x : EReal) : n • x = n * x := by + induction n with + | zero => rw [zero_smul, Nat.cast_zero, zero_mul] + | succ n ih => + rw [succ_nsmul, ih, Nat.cast_succ] + convert (EReal.right_distrib_of_nonneg _ _).symm <;> simp + /-! ### Absolute value -/ -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11215): TODO: use `Real.nnabs` for the case `(x : ℝ)` @@ -1651,7 +1852,7 @@ lemma antitone_div_right_of_nonpos {b : EReal} (h : b ≤ 0) : Antitone fun a change a' * b⁻¹ ≤ a * b⁻¹ rw [← neg_neg (a * b⁻¹), ← neg_neg (a' * b⁻¹), neg_le_neg_iff, mul_comm a b⁻¹, mul_comm a' b⁻¹, ← neg_mul b⁻¹ a, ← neg_mul b⁻¹ a', mul_comm (-b⁻¹) a, mul_comm (-b⁻¹) a', ← inv_neg b] - have : 0 ≤ -b := by apply le_neg_of_le_neg; simp [h] + have : 0 ≤ -b := by apply EReal.le_neg_of_le_neg; simp [h] exact div_le_div_right_of_nonneg this h' lemma div_le_div_right_of_nonpos {a a' b : EReal} (h : b ≤ 0) (h' : a ≤ a') : @@ -1745,4 +1946,4 @@ unsafe def positivity_coe_ennreal_ereal : expr → tactic strictness end Tactic -/ -set_option linter.style.longFile 1800 +set_option linter.style.longFile 2000 diff --git a/Mathlib/Topology/Instances/EReal.lean b/Mathlib/Topology/Instances/EReal.lean index 3af0b2ab4b18d..bb99c10fb4473 100644 --- a/Mathlib/Topology/Instances/EReal.lean +++ b/Mathlib/Topology/Instances/EReal.lean @@ -279,7 +279,7 @@ theorem continuousAt_add_top_coe (a : ℝ) : simp only [ContinuousAt, tendsto_nhds_top_iff_real, top_add_coe] refine fun r ↦ ((lt_mem_nhds (coe_lt_top (r - (a - 1)))).prod_nhds (lt_mem_nhds <| EReal.coe_lt_coe_iff.2 <| sub_one_lt _)).mono fun _ h ↦ ?_ - simpa only [← coe_add, sub_add_cancel] using add_lt_add h.1 h.2 + simpa only [← coe_add, _root_.sub_add_cancel] using add_lt_add h.1 h.2 theorem continuousAt_add_coe_top (a : ℝ) : ContinuousAt (fun p : EReal × EReal => p.1 + p.2) (a, ⊤) := by @@ -297,7 +297,7 @@ theorem continuousAt_add_bot_coe (a : ℝ) : simp only [ContinuousAt, tendsto_nhds_bot_iff_real, bot_add] refine fun r ↦ ((gt_mem_nhds (bot_lt_coe (r - (a + 1)))).prod_nhds (gt_mem_nhds <| EReal.coe_lt_coe_iff.2 <| lt_add_one _)).mono fun _ h ↦ ?_ - simpa only [← coe_add, sub_add_cancel] using add_lt_add h.1 h.2 + simpa only [← coe_add, _root_.sub_add_cancel] using add_lt_add h.1 h.2 theorem continuousAt_add_coe_bot (a : ℝ) : ContinuousAt (fun p : EReal × EReal => p.1 + p.2) (a, ⊥) := by From 88c8ffa1206368f30e35ee37dbd7157c5a1af810 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Tue, 17 Dec 2024 16:57:40 +0000 Subject: [PATCH 756/829] feat: pointwise MapsTo lemmas (#19999) --- .../Algebra/Group/Pointwise/Set/Basic.lean | 26 +++++++++++++++++++ .../Combinatorics/Additive/FreimanHom.lean | 9 +++---- 2 files changed, 29 insertions(+), 6 deletions(-) diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean index 641e71ad326f6..2582ba633719e 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean @@ -1463,9 +1463,35 @@ lemma preimage_div (hm : Injective m) {s t : Set β} (hs : s ⊆ range m) (ht : end Group +section Pi + variable {ι : Type*} {α : ι → Type*} [∀ i, Inv (α i)] @[to_additive (attr := simp)] lemma inv_pi (s : Set ι) (t : ∀ i, Set (α i)) : (s.pi t)⁻¹ = s.pi fun i ↦ (t i)⁻¹ := by ext x; simp +end Pi + +section Pointwise + +open scoped Pointwise + +@[to_additive] +lemma MapsTo.mul [Mul β] {A : Set α} {B₁ B₂ : Set β} {f₁ f₂ : α → β} + (h₁ : MapsTo f₁ A B₁) (h₂ : MapsTo f₂ A B₂) : MapsTo (f₁ * f₂) A (B₁ * B₂) := + fun _ h => mul_mem_mul (h₁ h) (h₂ h) + +@[to_additive] +lemma MapsTo.inv [DivisionCommMonoid β] {A : Set α} {B : Set β} {f : α → β} (h : MapsTo f A B) : + MapsTo (f⁻¹) A (B⁻¹) := + fun _ ha => inv_mem_inv.2 (h ha) + + +@[to_additive] +lemma MapsTo.div [DivisionCommMonoid β] {A : Set α} {B₁ B₂ : Set β} {f₁ f₂ : α → β} + (h₁ : MapsTo f₁ A B₁) (h₂ : MapsTo f₂ A B₂) : MapsTo (f₁ / f₂) A (B₁ / B₂) := + fun _ ha => div_mem_div (h₁ ha) (h₂ ha) + +end Pointwise + end Set diff --git a/Mathlib/Combinatorics/Additive/FreimanHom.lean b/Mathlib/Combinatorics/Additive/FreimanHom.lean index b3ec72a21f901..bcd9d5b6fb774 100644 --- a/Mathlib/Combinatorics/Additive/FreimanHom.lean +++ b/Mathlib/Combinatorics/Additive/FreimanHom.lean @@ -234,8 +234,7 @@ lemma isMulFreimanIso_empty : IsMulFreimanIso n (∅ : Set α) (∅ : Set β) f @[to_additive] lemma IsMulFreimanHom.mul (h₁ : IsMulFreimanHom n A B₁ f₁) (h₂ : IsMulFreimanHom n A B₂ f₂) : IsMulFreimanHom n A (B₁ * B₂) (f₁ * f₂) where - -- TODO: Extract `Set.MapsTo.mul` from this proof - mapsTo _ ha := mul_mem_mul (h₁.mapsTo ha) (h₂.mapsTo ha) + mapsTo := h₁.mapsTo.mul h₂.mapsTo map_prod_eq_map_prod s t hsA htA hs ht h := by rw [Pi.mul_def, prod_map_mul, prod_map_mul, h₁.map_prod_eq_map_prod hsA htA hs ht h, h₂.map_prod_eq_map_prod hsA htA hs ht h] @@ -327,16 +326,14 @@ variable [CommMonoid α] [DivisionCommMonoid β] {A : Set α} {B : Set β} {f : @[to_additive] lemma IsMulFreimanHom.inv (hf : IsMulFreimanHom n A B f) : IsMulFreimanHom n A B⁻¹ f⁻¹ where - -- TODO: Extract `Set.MapsTo.inv` from this proof - mapsTo _ ha := inv_mem_inv.2 (hf.mapsTo ha) + mapsTo := hf.mapsTo.inv map_prod_eq_map_prod s t hsA htA hs ht h := by rw [Pi.inv_def, prod_map_inv, prod_map_inv, hf.map_prod_eq_map_prod hsA htA hs ht h] @[to_additive] lemma IsMulFreimanHom.div {β : Type*} [DivisionCommMonoid β] {B₁ B₂ : Set β} {f₁ f₂ : α → β} (h₁ : IsMulFreimanHom n A B₁ f₁) (h₂ : IsMulFreimanHom n A B₂ f₂) : IsMulFreimanHom n A (B₁ / B₂) (f₁ / f₂) where - -- TODO: Extract `Set.MapsTo.div` from this proof - mapsTo _ ha := div_mem_div (h₁.mapsTo ha) (h₂.mapsTo ha) + mapsTo := h₁.mapsTo.div h₂.mapsTo map_prod_eq_map_prod s t hsA htA hs ht h := by rw [Pi.div_def, prod_map_div, prod_map_div, h₁.map_prod_eq_map_prod hsA htA hs ht h, h₂.map_prod_eq_map_prod hsA htA hs ht h] From 9890d5ff187e6a1625705cce1f16623a755a22db Mon Sep 17 00:00:00 2001 From: Xavier Roblot Date: Tue, 17 Dec 2024 18:11:53 +0000 Subject: [PATCH 757/829] feat(Analysis/PolarCoord): add versions for Lebesgue integral (#18490) This PR is part of the proof of the Analytic Class Number Formula. --- .../Analysis/SpecialFunctions/PolarCoord.lean | 61 +++++++++++++------ 1 file changed, 43 insertions(+), 18 deletions(-) diff --git a/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean b/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean index 46513cef6be86..016678accf824 100644 --- a/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean +++ b/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean @@ -22,7 +22,7 @@ noncomputable section Real open Real Set MeasureTheory -open scoped Real Topology +open scoped ENNReal Real Topology /-- The polar coordinates partial homeomorphism in `ℝ^2`, mapping `(r cos θ, r sin θ)` to `(r, θ)`. It is a homeomorphism between `ℝ^2 - (-∞, 0]` and `(0, +∞) × (-π, π)`. -/ @@ -99,6 +99,14 @@ theorem hasFDerivAt_polarCoord_symm (p : ℝ × ℝ) : (hasFDerivAt_fst.mul ((hasDerivAt_sin p.2).comp_hasFDerivAt p hasFDerivAt_snd)) using 2 <;> simp [smul_smul, add_comm, neg_mul, smul_neg, neg_smul _ (ContinuousLinearMap.snd ℝ ℝ ℝ)] +theorem det_fderiv_polarCoord_symm (p : ℝ × ℝ) : + (LinearMap.toContinuousLinearMap (Matrix.toLin (Basis.finTwoProd ℝ) (Basis.finTwoProd ℝ) + !![cos p.2, -p.1 * sin p.2; sin p.2, p.1 * cos p.2])).det = p.1 := by + conv_rhs => rw [← one_mul p.1, ← cos_sq_add_sin_sq p.2] + simp only [neg_mul, LinearMap.det_toContinuousLinearMap, LinearMap.det_toLin, + Matrix.det_fin_two_of, sub_neg_eq_add] + ring + -- Porting note: this instance is needed but not automatically synthesised instance : Measure.IsAddHaarMeasure volume (G := ℝ × ℝ) := Measure.prod.instIsAddHaarMeasure _ _ @@ -121,29 +129,36 @@ theorem polarCoord_source_ae_eq_univ : polarCoord.source =ᵐ[volume] univ := by theorem integral_comp_polarCoord_symm {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] (f : ℝ × ℝ → E) : (∫ p in polarCoord.target, p.1 • f (polarCoord.symm p)) = ∫ p, f p := by - set B : ℝ × ℝ → ℝ × ℝ →L[ℝ] ℝ × ℝ := fun p => - LinearMap.toContinuousLinearMap (Matrix.toLin (Basis.finTwoProd ℝ) (Basis.finTwoProd ℝ) - !![cos p.2, -p.1 * sin p.2; sin p.2, p.1 * cos p.2]) - have A : ∀ p ∈ polarCoord.symm.source, HasFDerivAt polarCoord.symm (B p) p := fun p _ => - hasFDerivAt_polarCoord_symm p - have B_det : ∀ p, (B p).det = p.1 := by - intro p - conv_rhs => rw [← one_mul p.1, ← cos_sq_add_sin_sq p.2] - simp only [B, neg_mul, LinearMap.det_toContinuousLinearMap, LinearMap.det_toLin, - Matrix.det_fin_two_of, sub_neg_eq_add] - ring symm calc ∫ p, f p = ∫ p in polarCoord.source, f p := by rw [← setIntegral_univ] apply setIntegral_congr_set exact polarCoord_source_ae_eq_univ.symm - _ = ∫ p in polarCoord.target, abs (B p).det • f (polarCoord.symm p) := by - apply integral_target_eq_integral_abs_det_fderiv_smul volume A + _ = ∫ p in polarCoord.target, |p.1| • f (polarCoord.symm p) := by + rw [← PartialHomeomorph.symm_target, integral_target_eq_integral_abs_det_fderiv_smul volume + (fun p _ ↦ hasFDerivAt_polarCoord_symm p), PartialHomeomorph.symm_source] + simp_rw [det_fderiv_polarCoord_symm] _ = ∫ p in polarCoord.target, p.1 • f (polarCoord.symm p) := by apply setIntegral_congr_fun polarCoord.open_target.measurableSet fun x hx => ?_ - rw [B_det, abs_of_pos] - exact hx.1 + rw [abs_of_pos hx.1] + +theorem lintegral_comp_polarCoord_symm (f : ℝ × ℝ → ℝ≥0∞) : + ∫⁻ (p : ℝ × ℝ) in polarCoord.target, ENNReal.ofReal p.1 • f (polarCoord.symm p) = + ∫⁻ (p : ℝ × ℝ), f p := by + symm + calc + _ = ∫⁻ p in polarCoord.symm '' polarCoord.target, f p := by + rw [← setLIntegral_univ, setLIntegral_congr polarCoord_source_ae_eq_univ.symm, + polarCoord.symm_image_target_eq_source ] + _ = ∫⁻ (p : ℝ × ℝ) in polarCoord.target, ENNReal.ofReal |p.1| • f (polarCoord.symm p) := by + rw [lintegral_image_eq_lintegral_abs_det_fderiv_mul volume _ + (fun p _ ↦ (hasFDerivAt_polarCoord_symm p).hasFDerivWithinAt)] + · simp_rw [det_fderiv_polarCoord_symm]; rfl + exacts [polarCoord.symm.injOn, measurableSet_Ioi.prod measurableSet_Ioo] + _ = ∫⁻ (p : ℝ × ℝ) in polarCoord.target, ENNReal.ofReal p.1 • f (polarCoord.symm p) := by + refine setLIntegral_congr_fun polarCoord.open_target.measurableSet ?_ + filter_upwards with _ hx using by rw [abs_of_pos hx.1] end Real @@ -151,7 +166,7 @@ noncomputable section Complex namespace Complex -open scoped Real +open scoped Real ENNReal /-- The polar coordinates partial homeomorphism in `ℂ`, mapping `r (cos θ + I * sin θ)` to `(r, θ)`. It is a homeomorphism between `ℂ - ℝ≤0` and `(0, +∞) × (-π, π)`. -/ @@ -173,6 +188,9 @@ protected theorem polarCoord_symm_apply (p : ℝ × ℝ) : Complex.polarCoord.symm p = p.1 * (Real.cos p.2 + Real.sin p.2 * Complex.I) := by simp [Complex.polarCoord, equivRealProdCLM_symm_apply, mul_add, mul_assoc] +theorem measurableEquivRealProd_symm_polarCoord_symm_apply (p : ℝ × ℝ) : + (measurableEquivRealProd.symm (polarCoord.symm p)) = Complex.polarCoord.symm p := rfl + theorem polarCoord_symm_abs (p : ℝ × ℝ) : Complex.abs (Complex.polarCoord.symm p) = |p.1| := by simp @@ -183,6 +201,13 @@ protected theorem integral_comp_polarCoord_symm {E : Type*} [NormedAddCommGroup (∫ p in polarCoord.target, p.1 • f (Complex.polarCoord.symm p)) = ∫ p, f p := by rw [← (Complex.volume_preserving_equiv_real_prod.symm).integral_comp measurableEquivRealProd.symm.measurableEmbedding, ← integral_comp_polarCoord_symm] - rfl + simp_rw [measurableEquivRealProd_symm_polarCoord_symm_apply] + +protected theorem lintegral_comp_polarCoord_symm (f : ℂ → ℝ≥0∞) : + (∫⁻ p in polarCoord.target, ENNReal.ofReal p.1 • f (Complex.polarCoord.symm p)) = + ∫⁻ p, f p := by + rw [← (volume_preserving_equiv_real_prod.symm).lintegral_comp_emb + measurableEquivRealProd.symm.measurableEmbedding, ← lintegral_comp_polarCoord_symm] + simp_rw [measurableEquivRealProd_symm_polarCoord_symm_apply] end Complex From 9282f9eb8423e3ef68a61e0a13d1129179d11821 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Tue, 17 Dec 2024 21:17:17 +0000 Subject: [PATCH 758/829] feat(Analysis/Fourier): Multi-variable Fourier series (#19766) --- Mathlib.lean | 1 + Mathlib/Analysis/Fourier/AddCircleMulti.lean | 264 +++++++++++++++++++ 2 files changed, 265 insertions(+) create mode 100644 Mathlib/Analysis/Fourier/AddCircleMulti.lean diff --git a/Mathlib.lean b/Mathlib.lean index 032b33531b73b..a5829df84a252 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1315,6 +1315,7 @@ import Mathlib.Analysis.Distribution.AEEqOfIntegralContDiff import Mathlib.Analysis.Distribution.FourierSchwartz import Mathlib.Analysis.Distribution.SchwartzSpace import Mathlib.Analysis.Fourier.AddCircle +import Mathlib.Analysis.Fourier.AddCircleMulti import Mathlib.Analysis.Fourier.FiniteAbelian.Orthogonality import Mathlib.Analysis.Fourier.FiniteAbelian.PontryaginDuality import Mathlib.Analysis.Fourier.FourierTransform diff --git a/Mathlib/Analysis/Fourier/AddCircleMulti.lean b/Mathlib/Analysis/Fourier/AddCircleMulti.lean new file mode 100644 index 0000000000000..6073814a5022b --- /dev/null +++ b/Mathlib/Analysis/Fourier/AddCircleMulti.lean @@ -0,0 +1,264 @@ +/- +Copyright (c) 2023 David Loeffler. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Loeffler +-/ +import Mathlib.Analysis.Fourier.AddCircle +import Mathlib.MeasureTheory.Integral.Pi + +/-! +# Multivariate Fourier series + +In this file we define the Fourier series of an L² function on the `d`-dimensional unit circle, and +show that it converges to the function in the L² norm. We also prove uniform convergence of the +Fourier series if `f` is continuous and the sequence of its Fourier coefficients is summable. +-/ + +noncomputable section + +open scoped BigOperators ComplexConjugate ENNReal + +open Set Algebra Submodule MeasureTheory + +-- some instances for unit circle + +attribute [local instance] Real.fact_zero_lt_one + +/-- In this file we normalise the measure on `ℝ / ℤ` to have total volume 1. -/ +local instance : MeasureSpace UnitAddCircle := ⟨AddCircle.haarAddCircle⟩ + +/-- The measure on `ℝ / ℤ` is a Haar measure. -/ +local instance : Measure.IsAddHaarMeasure (volume : Measure UnitAddCircle) := + inferInstanceAs (Measure.IsAddHaarMeasure AddCircle.haarAddCircle) + +/-- The measure on `ℝ / ℤ` is a probability measure. -/ +local instance : IsProbabilityMeasure (volume : Measure UnitAddCircle) := + inferInstanceAs (IsProbabilityMeasure AddCircle.haarAddCircle) + +/-- The product of finitely many copies of the unit circle, indexed by `d`. -/ +abbrev UnitAddTorus (d : Type*) := d → UnitAddCircle + +namespace UnitAddTorus + +variable {d : Type*} [Fintype d] + +section Monomials + +variable (n : d → ℤ) + +/-- Exponential monomials in `d` variables. -/ +def mFourier : C(UnitAddTorus d, ℂ) where + toFun x := ∏ i : d, fourier (n i) (x i) + continuous_toFun := continuous_finset_prod _ + fun i _ ↦ (fourier (n i)).continuous.comp (continuous_apply i) + +variable {n} {x : UnitAddTorus d} + +lemma mFourier_neg : mFourier (-n) x = conj (mFourier n x) := by + simp only [mFourier, Pi.neg_apply, fourier_neg, ContinuousMap.coe_mk, map_prod] + +lemma mFourier_add {m : d → ℤ} : mFourier (m + n) x = mFourier m x * mFourier n x := by + simp only [mFourier, Pi.add_apply, fourier_add, ContinuousMap.coe_mk, ← Finset.prod_mul_distrib] + +lemma mFourier_zero : mFourier (0 : d → ℤ) = 1 := by + ext x + simp only [mFourier, Pi.zero_apply, fourier_zero, Finset.prod_const_one, ContinuousMap.coe_mk, + ContinuousMap.one_apply] + +lemma mFourier_norm : ‖mFourier n‖ = 1 := by + apply le_antisymm + · refine (ContinuousMap.norm_le _ zero_le_one).mpr fun i ↦ ?_ + simp only [mFourier, fourier_apply, ContinuousMap.coe_mk, norm_prod, Complex.norm_eq_abs, + Circle.abs_coe, Finset.prod_const_one, le_rfl] + · refine (le_of_eq ?_).trans ((mFourier n).norm_coe_le_norm fun _ ↦ 0) + simp only [mFourier, ContinuousMap.coe_mk, fourier_eval_zero, Finset.prod_const_one, + CStarRing.norm_one] + +lemma mFourier_single [DecidableEq d] (z : d → AddCircle (1 : ℝ)) (i : d) : + mFourier (Pi.single i 1) z = fourier 1 (z i) := by + simp_rw [mFourier, ContinuousMap.coe_mk] + have := Finset.prod_mul_prod_compl {i} (fun j ↦ fourier ((Pi.single i (1 : ℤ) : d → ℤ) j) (z j)) + rw [Finset.prod_singleton, Finset.prod_congr rfl (fun j hj ↦ ?_)] at this + · rw [← this, Finset.prod_const_one, mul_one, Pi.single_eq_same] + · rw [Finset.mem_compl, Finset.mem_singleton] at hj + simp only [Pi.single_eq_of_ne hj, fourier_zero] + +end Monomials + +section Algebra + +/-- The star subalgebra of `C(UnitAddTorus d, ℂ)` generated by `mFourier n` for `n ∈ ℤᵈ`. -/ +def mFourierSubalgebra (d : Type*) [Fintype d] : StarSubalgebra ℂ C(UnitAddTorus d, ℂ) where + toSubalgebra := Algebra.adjoin ℂ (range mFourier) + star_mem' := by + show Algebra.adjoin ℂ (range mFourier) ≤ star (Algebra.adjoin ℂ (range mFourier)) + refine adjoin_le ?_ + rintro _ ⟨n, rfl⟩ + refine subset_adjoin ⟨-n, ?_⟩ + ext1 x + simp only [mFourier_neg, starRingEnd_apply, ContinuousMap.star_apply] + +/-- The star subalgebra of `C(UnitAddTorus d, ℂ)` generated by `mFourier n` for `n ∈ ℤᵈ` is in fact +the linear span of these functions. -/ +theorem mFourierSubalgebra_coe : + (mFourierSubalgebra d).toSubalgebra.toSubmodule = span ℂ (range mFourier) := by + apply adjoin_eq_span_of_subset + refine .trans (fun x ↦ Submonoid.closure_induction (fun _ ↦ id) ⟨0, ?_⟩ ?_) subset_span + · ext z + simp only [mFourier, Pi.zero_apply, fourier_zero, Finset.prod_const, one_pow, + ContinuousMap.coe_mk, ContinuousMap.one_apply] + · rintro _ _ _ _ ⟨m, rfl⟩ ⟨n, rfl⟩ + refine ⟨m + n, ?_⟩ + ext z + simp only [mFourier, Pi.add_apply, fourier_apply, fourier_add', Finset.prod_mul_distrib, + ContinuousMap.coe_mk, ContinuousMap.mul_apply] + +/-- The subalgebra of `C(UnitAddTorus d, ℂ)` generated by `mFourier n` for `n ∈ ℤᵈ` separates +points. -/ +theorem mFourierSubalgebra_separatesPoints : (mFourierSubalgebra d).SeparatesPoints := by + classical + intro x y hxy + rw [Ne, funext_iff, not_forall] at hxy + obtain ⟨i, hi⟩ := hxy + refine ⟨_, ⟨mFourier (Pi.single i 1), subset_adjoin ⟨Pi.single i 1, rfl⟩, rfl⟩, ?_⟩ + dsimp only + rw [mFourier_single, mFourier_single, fourier_one, fourier_one, Ne, Subtype.coe_inj] + contrapose! hi + exact AddCircle.injective_toCircle one_ne_zero hi + +/-- The subalgebra of `C(UnitAddTorus d, ℂ)` generated by `mFourier n` for `n : d → ℤ` is dense. -/ +theorem mFourierSubalgebra_closure_eq_top : (mFourierSubalgebra d).topologicalClosure = ⊤ := + ContinuousMap.starSubalgebra_topologicalClosure_eq_top_of_separatesPoints _ + mFourierSubalgebra_separatesPoints + +/-- The linear span of the monomials `mFourier n` is dense in `C(UnitAddTorus d, ℂ)`. -/ +theorem span_mFourier_closure_eq_top : + (span ℂ (range <| mFourier (d := d))).topologicalClosure = ⊤ := by + rw [← mFourierSubalgebra_coe] + exact congr_arg (Subalgebra.toSubmodule <| StarSubalgebra.toSubalgebra ·) + mFourierSubalgebra_closure_eq_top + +end Algebra + +section Lp + +/-- The family of monomials `mFourier n`, parametrized by `n : ℤᵈ` and considered as +elements of the `Lp` space of functions `UnitAddTorus d → ℂ`. -/ +abbrev mFourierLp (p : ℝ≥0∞) [Fact (1 ≤ p)] (n : d → ℤ) : + Lp ℂ p (volume : Measure (UnitAddTorus d)) := + ContinuousMap.toLp (E := ℂ) p volume ℂ (mFourier n) + +theorem coeFn_mFourierLp (p : ℝ≥0∞) [Fact (1 ≤ p)] (n : d → ℤ) : + mFourierLp p n =ᵐ[volume] mFourier n := + ContinuousMap.coeFn_toLp volume (mFourier n) + +/-- For each `1 ≤ p < ∞`, the linear span of the monomials `mFourier n` is dense in the `Lᵖ` space +of functions on `UnitAddTorus d`. -/ +theorem span_mFourierLp_closure_eq_top {p : ℝ≥0∞} [Fact (1 ≤ p)] (hp : p ≠ ∞) : + (span ℂ (range (@mFourierLp d _ p _))).topologicalClosure = ⊤ := by + simpa only [map_span, ContinuousLinearMap.coe_coe, ← range_comp, Function.comp_def] using + (ContinuousMap.toLp_denseRange ℂ volume ℂ hp).topologicalClosure_map_submodule + span_mFourier_closure_eq_top + +/-- The monomials `mFourierLp 2 n` are an orthonormal set in `L²`. -/ +theorem orthonormal_mFourier : Orthonormal ℂ (mFourierLp (d := d) 2) := by + rw [orthonormal_iff_ite] + intro m n + simp only [ContinuousMap.inner_toLp, ← mFourier_neg, ← mFourier_add] + split_ifs with h + · simpa only [h, neg_add_cancel, mFourier_zero, measure_univ, ENNReal.one_toReal, one_smul] using + integral_const (α := UnitAddTorus d) (μ := volume) (1 : ℂ) + rw [mFourier, ContinuousMap.coe_mk, MeasureTheory.integral_fintype_prod_eq_prod] + obtain ⟨i, hi⟩ := Function.ne_iff.mp h + apply Finset.prod_eq_zero (Finset.mem_univ i) + simpa only [eq_false_intro hi, if_false, ContinuousMap.inner_toLp, ← fourier_neg, + ← fourier_add] using (orthonormal_iff_ite.mp <| orthonormal_fourier) (m i) (n i) + +end Lp + +section fourierCoeff + +variable {E : Type} [NormedAddCommGroup E] [NormedSpace ℂ E] + +/-- The `n`-th Fourier coefficient of a function `UnitAddTorus d → E`, for `E` a complete normed +`ℂ`-vector space, defined as the integral over `UnitAddTorus d` of `mFourier (-n) t • f t`. -/ +def mFourierCoeff (f : UnitAddTorus d → E) (n : d → ℤ) : E := ∫ t, mFourier (-n) t • f t + +end fourierCoeff + +section FourierL2 + +local notation "L²(" α ")" => Lp ℂ 2 (volume : Measure α) + +/-- We define `mFourierBasis` to be a `ℤᵈ`-indexed Hilbert basis for the `L²` space of functions +on `UnitAddTorus d`, which by definition is an isometric isomorphism from `L²(UnitAddTorus d)` +to `ℓ²(ℤᵈ, ℂ)`. -/ +def mFourierBasis : HilbertBasis (d → ℤ) ℂ L²(UnitAddTorus d) := + HilbertBasis.mk orthonormal_mFourier (span_mFourierLp_closure_eq_top (by norm_num)).ge + +/-- The elements of the Hilbert basis `mFourierBasis` are the functions `mFourierLp 2`, i.e. the +monomials `mFourier n` on `UnitAddTorus d` considered as elements of `L²`. -/ +@[simp] +theorem coe_mFourierBasis : ⇑(mFourierBasis (d := d)) = mFourierLp 2 := HilbertBasis.coe_mk _ _ + +/-- Under the isometric isomorphism `mFourierBasis` from `L²(UnitAddTorus d)` to `ℓ²(ℤᵈ, ℂ)`, +the `i`-th coefficient is `mFourierCoeff f i`. -/ +theorem mFourierBasis_repr (f : L²(UnitAddTorus d)) (i : d → ℤ) : + mFourierBasis.repr f i = mFourierCoeff f i := by + trans ∫ t, conj (mFourierLp 2 i t) * f t + · rw [mFourierBasis.repr_apply_apply f i, MeasureTheory.L2.inner_def, coe_mFourierBasis] + simp only [RCLike.inner_apply] + · apply integral_congr_ae + filter_upwards [coeFn_mFourierLp 2 i] with _ ht + rw [ht, ← mFourier_neg, smul_eq_mul] + +/-- The Fourier series of an `L2` function `f` sums to `f` in the `L²` norm. -/ +theorem hasSum_mFourier_series_L2 (f : L²(UnitAddTorus d)) : + HasSum (fun i ↦ mFourierCoeff f i • mFourierLp 2 i) f := by + simpa [← coe_mFourierBasis, mFourierBasis_repr] using mFourierBasis.hasSum_repr f + +/-- **Parseval's identity** for inner products: for `L²` functions `f, g` on `UnitAddTorus d`, the +inner product of the Fourier coefficients of `f` and `g` is the inner product of `f` and `g`. -/ +theorem hasSum_prod_mFourierCoeff (f g : L²(UnitAddTorus d)) : + HasSum (fun i ↦ conj (mFourierCoeff f i) * (mFourierCoeff g i)) (∫ t, conj (f t) * g t) := by + refine HasSum.congr_fun (mFourierBasis.hasSum_inner_mul_inner f g) (fun n ↦ ?_) + simp only [← mFourierBasis_repr, HilbertBasis.repr_apply_apply, inner_conj_symm] + +/-- **Parseval's identity** for norms: for an `L²` function `f` on `UnitAddTorus d`, the sum of the +squared norms of the Fourier coefficients equals the `L²` norm of `f`. -/ +theorem hasSum_sq_mFourierCoeff (f : L²(UnitAddTorus d)) : + HasSum (fun i ↦ ‖mFourierCoeff f i‖ ^ 2) (∫ t, ‖f t‖ ^ 2) := by + simpa only [← RCLike.inner_apply, inner_self_eq_norm_sq, ← integral_re + (L2.integrable_inner f f)] using RCLike.hasSum_re ℂ (hasSum_prod_mFourierCoeff f f) + +end FourierL2 + +section Convergence + +variable (f : C(UnitAddTorus d, ℂ)) + +theorem mFourierCoeff_toLp (n : d → ℤ) : + mFourierCoeff (f.toLp 2 volume ℂ) n = mFourierCoeff f n := + integral_congr_ae (ae_eq_rfl.mul <| f.coeFn_toAEEqFun _) + +variable {f} + +/-- If the sequence of Fourier coefficients of `f` is summable, then the Fourier series converges +uniformly to `f`. -/ +theorem hasSum_mFourier_series_of_summable (h : Summable (mFourierCoeff f)) : + HasSum (fun i ↦ mFourierCoeff f i • mFourier i) f := by + have sum_L2 := hasSum_mFourier_series_L2 (ContinuousMap.toLp 2 volume ℂ f) + simp only [mFourierCoeff_toLp] at sum_L2 + refine ContinuousMap.hasSum_of_hasSum_Lp (.of_norm ?_) sum_L2 + simpa only [norm_smul, mFourier_norm, mul_one] using h.norm + +/-- If the sequence of Fourier coefficients of `f` is summable, then the Fourier series of `f` +converges everywhere pointwise to `f`. -/ +theorem hasSum_mFourier_series_apply_of_summable (h : Summable (mFourierCoeff f)) + (x : UnitAddTorus d) : HasSum (fun i ↦ mFourierCoeff f i • mFourier i x) (f x) := by + simpa only [_root_.map_smul] using (ContinuousMap.evalCLM ℂ x).hasSum + (hasSum_mFourier_series_of_summable h) + +end Convergence + +end UnitAddTorus From 495ec69bfbbe8fe20099db76e889e39b676ba43f Mon Sep 17 00:00:00 2001 From: David Renshaw Date: Tue, 17 Dec 2024 21:44:13 +0000 Subject: [PATCH 759/829] fix(Linter/DocPrime): don't flag anonymous examples and instances (#20027) Fixes a bug reported [in Zulip](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/strange.20docPrime.20linter.20behaviour/near/489550515). When an `example` or anonymous `instance` occurs within a primed namespace, the `docPrime` linter should *not* report a message. --- Mathlib/Tactic/Linter/DocPrime.lean | 3 +++ MathlibTest/DocPrime.lean | 8 ++++++++ 2 files changed, 11 insertions(+) diff --git a/Mathlib/Tactic/Linter/DocPrime.lean b/Mathlib/Tactic/Linter/DocPrime.lean index c2a8a38fbcdcf..2ccd4de91802f 100644 --- a/Mathlib/Tactic/Linter/DocPrime.lean +++ b/Mathlib/Tactic/Linter/DocPrime.lean @@ -45,6 +45,8 @@ def docPrimeLinter : Linter where run := withSetOptionIn fun stx ↦ do 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 + -- ignore examples + if (stx.find? (·.isOfKind ``Lean.Parser.Command.example)).isSome then return let docstring := stx[0][0] -- The current declaration's id, possibly followed by a list of universe names. let declId := @@ -52,6 +54,7 @@ def docPrimeLinter : Linter where run := withSetOptionIn fun stx ↦ do stx[1][3][0] else stx[1][1] + if let .missing := declId then return -- The name of the current declaration, with namespaces resolved. let declName : Name := if let `_root_ :: rest := declId[0].getId.components then diff --git a/MathlibTest/DocPrime.lean b/MathlibTest/DocPrime.lean index 3809fa9675e65..86f26d1369df8 100644 --- a/MathlibTest/DocPrime.lean +++ b/MathlibTest/DocPrime.lean @@ -78,3 +78,11 @@ note: this linter can be disabled with `set_option linter.docPrime false` -/ #guard_msgs in def def_no_doc' : True := .intro + +-- Anonymous declarations in a primed namespace should not get flagged by the linter. +namespace Foo' + +example : True := .intro +instance : True := .intro + +end Foo' From fb40f7550d02605e9a7bf195b85184714c5d06dc Mon Sep 17 00:00:00 2001 From: Alex Brodbelt Date: Wed, 18 Dec 2024 01:10:20 +0000 Subject: [PATCH 760/829] feat(Archive/Imo): formalize IMO 1982q3 (#16190) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Formalize problem 3 of 1982 IMO. Co-authored by: Violeta Hernández Palacios Co-authored-by: AlexBrodbelt Co-authored-by: Alex Brodbelt <64128135+AlexBrodbelt@users.noreply.github.com> Co-authored-by: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> --- Archive.lean | 1 + Archive/Imo/Imo1982Q3.lean | 93 +++++++++++++++++++ Mathlib/Algebra/GeomSum.lean | 8 ++ .../Order/GroupWithZero/Unbundled.lean | 2 +- 4 files changed, 103 insertions(+), 1 deletion(-) create mode 100644 Archive/Imo/Imo1982Q3.lean diff --git a/Archive.lean b/Archive.lean index 3b1428fa31cb4..fc8da3829ecc9 100644 --- a/Archive.lean +++ b/Archive.lean @@ -20,6 +20,7 @@ import Archive.Imo.Imo1975Q1 import Archive.Imo.Imo1977Q6 import Archive.Imo.Imo1981Q3 import Archive.Imo.Imo1982Q1 +import Archive.Imo.Imo1982Q3 import Archive.Imo.Imo1986Q5 import Archive.Imo.Imo1987Q1 import Archive.Imo.Imo1988Q6 diff --git a/Archive/Imo/Imo1982Q3.lean b/Archive/Imo/Imo1982Q3.lean new file mode 100644 index 0000000000000..19a8833b152c6 --- /dev/null +++ b/Archive/Imo/Imo1982Q3.lean @@ -0,0 +1,93 @@ +/- +Copyright (c) 2024 Violeta Hernández Palacios. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Violeta Hernández Palacios, Alex Brodbelt +-/ +import Mathlib.Algebra.Order.BigOperators.Ring.Finset +import Mathlib.Algebra.GeomSum +import Mathlib.Data.NNReal.Basic + +/-! +# IMO 1982 Q3 + +Consider infinite sequences $\{x_n\}$ of positive reals such that $x_0 = 1$ and +$x_0 \ge x_1 \ge x_2 \ge \ldots$ + +a) Prove that for every such sequence there is an $n \ge 1$ such that: + +$$\frac{x_0^2}{x_1} + \ldots + \frac{x_{n-1}^2}{x_n} \ge 3.999$$ + +b) Find such a sequence such that for all $n$: + +$$\frac{x_0^2}{x_1} + \ldots + \frac{x_{n-1}^2}{x_n} < 4$$ + +The solution is based on Solution 1 from the +[Art of Problem Solving](https://artofproblemsolving.com/wiki/index.php/1982_IMO_Problems/Problem_3) +website. For part a, we use Sedrakyan's lemma to show the sum is bounded below by +$\frac{4n}{n + 1}$, which can be made arbitrarily close to $4$ by taking large $n$. For part b, we +show the sequence $x_n = 2^{-n}$ satisfies the desired inequality. +-/ + +open Finset NNReal + +variable {x : ℕ → ℝ} {n : ℕ} (hn : n ≠ 0) (hx : Antitone x) + +namespace Imo1982Q3 +include hn hx + +/-- `x n` is at most the average of all previous terms in the sequence. + +This is expressed here with `∑ k ∈ range n, x k` added to both sides. -/ +lemma le_avg : ∑ k ∈ range (n + 1), x k ≤ (∑ k ∈ range n, x k) * (1 + 1 / n) := by + rw [sum_range_succ, mul_one_add, add_le_add_iff_left, mul_one_div, + le_div_iff₀ (mod_cast hn.bot_lt), mul_comm, ← nsmul_eq_mul] + conv_lhs => rw [← card_range n, ← sum_const] + refine sum_le_sum fun k hk ↦ hx (le_of_lt ?_) + simpa using hk + +/-- The main inequality used for part a. -/ +lemma ineq (h0 : x 0 = 1) (hp : ∀ k, 0 < x k) : + 4 * n / (n + 1) ≤ ∑ k ∈ range (n + 1), x k ^ 2 / x (k + 1) := by + calc + -- We first use AM-GM. + _ ≤ (∑ k ∈ range n, x (k + 1) + 1) ^ 2 / (∑ k ∈ range n, x (k + 1)) * n / (n + 1) := by + gcongr + rw [le_div_iff₀] + · simpa using four_mul_le_sq_add (∑ k ∈ range n, x (k + 1)) 1 + · exact sum_pos (fun k _ ↦ hp _) (nonempty_range_iff.2 hn) + -- We move the fraction into the denominator. + _ = (∑ k ∈ range n, x (k + 1) + 1) ^ 2 / ((∑ k ∈ range n, x (k + 1)) * (1 + 1 / n)) := by + field_simp + -- We make use of the `le_avg` lemma. + _ ≤ (∑ k ∈ range (n + 1), x k) ^ 2 / ∑ k ∈ range (n + 1), x (k + 1) := by + gcongr + · exact sum_pos (fun k _ ↦ hp _) nonempty_range_succ + · exact add_nonneg (sum_nonneg fun k _ ↦ (hp _).le) zero_le_one + · rw [sum_range_succ', h0] + · exact le_avg hn (hx.comp_monotone @Nat.succ_le_succ) + -- We conclude by Sedrakyan. + _ ≤ _ := sq_sum_div_le_sum_sq_div _ x fun k _ ↦ hp (k + 1) + +end Imo1982Q3 + +/-- Part a of the problem is solved by `n = 4000`. -/ +theorem imo1982_q3a (hx : Antitone x) (h0 : x 0 = 1) (hp : ∀ k, 0 < x k) : + ∃ n : ℕ, 3.999 ≤ ∑ k ∈ range n, (x k) ^ 2 / x (k + 1) := by + use 4000 + convert Imo1982Q3.ineq (Nat.succ_ne_zero 3998) hx h0 hp + norm_num + +/-- Part b of the problem is solved by `x k = (1 / 2) ^ k`. -/ +theorem imo1982_q3b : ∃ x : ℕ → ℝ, Antitone x ∧ x 0 = 1 ∧ (∀ k, 0 < x k) + ∧ ∀ n, ∑ k ∈ range n, x k ^ 2 / x (k + 1) < 4 := by + refine ⟨fun k ↦ 2⁻¹ ^ k, ?_, pow_zero _, ?_, fun n ↦ ?_⟩ + · apply (pow_right_strictAnti₀ _ _).antitone <;> norm_num + · simp + · have {k : ℕ} : (2 : ℝ)⁻¹ ^ (k * 2) * ((2 : ℝ)⁻¹ ^ k)⁻¹ = (2 : ℝ)⁻¹ ^ k := by + rw [← pow_sub₀] <;> simp [mul_two] + simp_rw [← pow_mul, pow_succ, ← div_eq_mul_inv, div_div_eq_mul_div, mul_comm, mul_div_assoc, + ← mul_sum, div_eq_mul_inv, this, ← two_add_two_eq_four, ← mul_two, + mul_lt_mul_iff_of_pos_left two_pos] + convert NNReal.coe_lt_coe.2 <| geom_sum_lt (inv_ne_zero two_ne_zero) two_inv_lt_one n + · simp + · norm_num diff --git a/Mathlib/Algebra/GeomSum.lean b/Mathlib/Algebra/GeomSum.lean index 2a0df0564ca2b..f6736ac5b2e81 100644 --- a/Mathlib/Algebra/GeomSum.lean +++ b/Mathlib/Algebra/GeomSum.lean @@ -294,6 +294,14 @@ lemma geom_sum_of_lt_one {x : α} [CanonicallyLinearOrderedSemifield α] [Sub α ∑ i ∈ Finset.range n, x ^ i = (1 - x ^ n) / (1 - x) := eq_div_of_mul_eq (tsub_pos_of_lt h).ne' (geom_sum_mul_of_le_one h.le n) +theorem geom_sum_lt {x : α} [CanonicallyLinearOrderedSemifield α] [Sub α] [OrderedSub α] + (h0 : x ≠ 0) (h1 : x < 1) (n : ℕ) : ∑ i ∈ range n, x ^ i < (1 - x)⁻¹ := by + rw [← zero_lt_iff] at h0 + rw [geom_sum_of_lt_one h1, div_lt_iff₀, inv_mul_cancel₀, tsub_lt_self_iff] + · exact ⟨h0.trans h1, pow_pos h0 n⟩ + · rwa [ne_eq, tsub_eq_zero_iff_le, not_le] + · rwa [tsub_pos_iff_lt] + protected theorem Commute.mul_geom_sum₂_Ico [Ring α] {x y : α} (h : Commute x y) {m n : ℕ} (hmn : m ≤ n) : ((x - y) * ∑ i ∈ Finset.Ico m n, x ^ i * y ^ (n - 1 - i)) = x ^ n - x ^ m * y ^ (n - m) := by diff --git a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean index 2276d496cf83f..a9286e9934add 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean @@ -1637,7 +1637,7 @@ lemma mul_inv_lt_iff₀ (hc : 0 < c) : b * c⁻¹ < a ↔ b < a * c where 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. -/ +/-- See `div_lt_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] From d8e48a35237bfc90f26ca475fc198fe520a2e8e2 Mon Sep 17 00:00:00 2001 From: Arend Mellendijk Date: Wed, 18 Dec 2024 01:49:22 +0000 Subject: [PATCH 761/829] feat(BigOperators): prod_ite_eq_of_mem (#20026) I found these lemmas very helpful for manipulating sums in my Selberg sieve project. The benefit these lemmas have over `prod_ite_eq` and `prod_ite_eq'` is that the RHS is not an if-then-else block, which makes it much more useful for rewriting backwards to introduce a sum. --- Mathlib/Algebra/BigOperators/Group/Finset.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index 80ea8b0b9c9f1..ae179a61b91c4 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -1165,6 +1165,17 @@ theorem prod_ite_eq' [DecidableEq α] (s : Finset α) (a : α) (b : α → β) : (∏ x ∈ s, ite (x = a) (b x) 1) = ite (a ∈ s) (b a) 1 := prod_dite_eq' s a fun x _ => b x +@[to_additive] +theorem prod_ite_eq_of_mem [DecidableEq α] (s : Finset α) (a : α) (b : α → β) (h : a ∈ s) : + (∏ x ∈ s, if a = x then b x else 1) = b a := by + simp only [prod_ite_eq, if_pos h] + +/-- The difference with `Finset.prod_ite_eq_of_mem` is that the arguments to `Eq` are swapped. -/ +@[to_additive] +theorem prod_ite_eq_of_mem' [DecidableEq α] (s : Finset α) (a : α) (b : α → β) (h : a ∈ s) : + (∏ x ∈ s, if x = a then b x else 1) = b a := by + simp only [prod_ite_eq', if_pos h] + @[to_additive] theorem prod_ite_index (p : Prop) [Decidable p] (s t : Finset α) (f : α → β) : ∏ x ∈ if p then s else t, f x = if p then ∏ x ∈ s, f x else ∏ x ∈ t, f x := From 2144977fb121989a45d3f6e4d74fd8ab2350db3e Mon Sep 17 00:00:00 2001 From: Vasily Nesterov Date: Wed, 18 Dec 2024 06:37:04 +0000 Subject: [PATCH 762/829] feat(Analysis/Analytic): `FormalMultilinearSeries.unshift` convergence (#19848) * Prove `unshift_shift`, which states that `p.unshift.shift = p`. * Prove `radius_shift` and `radius_unshift`, which state that the shifted and unshifted series converge on the same ball. * Prove `HasFPowerSeriesOnBall.unshift` (and variations), describing where the shifted series converges. --- Mathlib/Analysis/Analytic/Basic.lean | 39 +++++++++++++++++ Mathlib/Analysis/Analytic/Constructions.lean | 42 +++++++++++++++++++ .../Calculus/FormalMultilinearSeries.lean | 6 +++ 3 files changed, 87 insertions(+) diff --git a/Mathlib/Analysis/Analytic/Basic.lean b/Mathlib/Analysis/Analytic/Basic.lean index e1e5359b0a1de..9cefdb7e129ca 100644 --- a/Mathlib/Analysis/Analytic/Basic.lean +++ b/Mathlib/Analysis/Analytic/Basic.lean @@ -327,6 +327,45 @@ theorem min_radius_le_radius_add (p q : FormalMultilinearSeries 𝕜 E F) : theorem radius_neg (p : FormalMultilinearSeries 𝕜 E F) : (-p).radius = p.radius := by simp only [radius, neg_apply, norm_neg] +@[simp] +theorem radius_shift (p : FormalMultilinearSeries 𝕜 E F) : p.shift.radius = p.radius := by + simp only [radius, shift, Nat.succ_eq_add_one, ContinuousMultilinearMap.curryRight_norm] + congr + ext r + apply eq_of_le_of_le + · apply iSup_mono' + intro C + use ‖p 0‖ ⊔ (C * r) + apply iSup_mono' + intro h + simp only [le_refl, le_sup_iff, exists_prop, and_true] + intro n + cases' n with m + · simp + right + rw [pow_succ, ← mul_assoc] + apply mul_le_mul_of_nonneg_right (h m) zero_le_coe + · apply iSup_mono' + intro C + use ‖p 1‖ ⊔ C / r + apply iSup_mono' + intro h + simp only [le_refl, le_sup_iff, exists_prop, and_true] + intro n + by_cases hr : r = 0 + · rw [hr] + cases n <;> simp + right + replace hr : 0 < (r : ℝ) := pos_iff_ne_zero.mpr hr + specialize h (n + 1) + rw [le_div_iff₀ hr] + rwa [pow_succ, ← mul_assoc] at h + +@[simp] +theorem radius_unshift (p : FormalMultilinearSeries 𝕜 E (E →L[𝕜] F)) (z : F) : + (p.unshift z).radius = p.radius := by + rw [← radius_shift, unshift_shift] + protected theorem hasSum [CompleteSpace F] (p : FormalMultilinearSeries 𝕜 E F) {x : E} (hx : x ∈ EMetric.ball (0 : E) p.radius) : HasSum (fun n : ℕ => p n fun _ => x) (p.sum x) := (p.summable hx).hasSum diff --git a/Mathlib/Analysis/Analytic/Constructions.lean b/Mathlib/Analysis/Analytic/Constructions.lean index c311af2a50540..f91da27484023 100644 --- a/Mathlib/Analysis/Analytic/Constructions.lean +++ b/Mathlib/Analysis/Analytic/Constructions.lean @@ -942,3 +942,45 @@ theorem Finset.analyticOnNhd_prod {A : Type*} [NormedCommRing A] [NormedAlgebra {f : α → E → A} {s : Set E} (N : Finset α) (h : ∀ n ∈ N, AnalyticOnNhd 𝕜 (f n) s) : AnalyticOnNhd 𝕜 (fun z ↦ ∏ n ∈ N, f n z) s := fun z zs ↦ N.analyticAt_prod (fun n m ↦ h n m z zs) + +/-! +### Unshifting +-/ + +section + +variable {f : E → (E →L[𝕜] F)} {pf : FormalMultilinearSeries 𝕜 E (E →L[𝕜] F)} {s : Set E} {x : E} + {r : ℝ≥0∞} {z : F} + +theorem HasFPowerSeriesWithinOnBall.unshift (hf : HasFPowerSeriesWithinOnBall f pf s x r) : + HasFPowerSeriesWithinOnBall (fun y ↦ z + f y (y - x)) (pf.unshift z) s x r where + r_le := by + rw [FormalMultilinearSeries.radius_unshift] + exact hf.r_le + r_pos := hf.r_pos + hasSum := by + intro y hy h'y + apply HasSum.zero_add + simp only [FormalMultilinearSeries.unshift, Nat.succ_eq_add_one, + continuousMultilinearCurryRightEquiv_symm_apply', add_sub_cancel_left] + exact (ContinuousLinearMap.apply 𝕜 F y).hasSum (hf.hasSum hy h'y) + +theorem HasFPowerSeriesOnBall.unshift (hf : HasFPowerSeriesOnBall f pf x r) : + HasFPowerSeriesOnBall (fun y ↦ z + f y (y - x)) (pf.unshift z) x r where + r_le := by + rw [FormalMultilinearSeries.radius_unshift] + exact hf.r_le + r_pos := hf.r_pos + hasSum := by + intro y hy + apply HasSum.zero_add + simp only [FormalMultilinearSeries.unshift, Nat.succ_eq_add_one, + continuousMultilinearCurryRightEquiv_symm_apply', add_sub_cancel_left] + exact (ContinuousLinearMap.apply 𝕜 F y).hasSum (hf.hasSum hy) + +theorem HasFPowerSeriesWithinAt.unshift (hf : HasFPowerSeriesWithinAt f pf s x) : + HasFPowerSeriesWithinAt (fun y ↦ z + f y (y - x)) (pf.unshift z) s x := + let ⟨_, hrf⟩ := hf + hrf.unshift.hasFPowerSeriesWithinAt + +end diff --git a/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean b/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean index d95f0f210244f..4b12e24cee9e5 100644 --- a/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean +++ b/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean @@ -179,6 +179,12 @@ def unshift (q : FormalMultilinearSeries 𝕜 E (E →L[𝕜] F)) (z : F) : Form | 0 => (continuousMultilinearCurryFin0 𝕜 E F).symm z | n + 1 => (continuousMultilinearCurryRightEquiv' 𝕜 n E F).symm (q n) +theorem unshift_shift {p : FormalMultilinearSeries 𝕜 E (E →L[𝕜] F)} {z : F} : + (p.unshift z).shift = p := by + ext1 n + simp [shift, unshift] + exact LinearIsometryEquiv.apply_symm_apply (continuousMultilinearCurryRightEquiv' 𝕜 n E F) (p n) + end FormalMultilinearSeries section From 2375546cb395e5d4e3a73428446fc4176155fa5f Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Wed, 18 Dec 2024 10:57:39 +0000 Subject: [PATCH 763/829] chore(Topology/Separation): split large file (#20031) This splits up the behemoth `Mathlib.Topology.Separation.Basic` (2700 lines) into 5 smaller files: - `Mathlib.Topology.Separation.SeparatedNhds`: auxiliary definitions of `SeparatedNhds` and `HasSeparatingCover` - `Mathlib.Topology.Separation.Basic`: T0, R0, T1 and R1 spaces - `Mathlib.Topology.Separation.Hausdorff`: T2 and T2.5 spaces - `Mathlib.Topology.Separation.Regular`: regular, normal, completely normal, T3-T5 - `Mathlib.Topology.Separation.Profinite`: totally disconnected compact Hausdorff spaces are totally separated. It also moves `Mathlib.Topology.CompletelyRegular` and `Mathlib.Topology.CountableSeparatingOn` into the `Separation` directory. The only changes outside the files listed above are updating `import` lines. The only change to actual code is in `Mathlib.Topology.Separation.Profinite`, where I realised that we had two theorems asserting exactly the same thing, one under strictly more general typeclass hypotheses than the other (locally compact rather than compact spaces). So I made the less general theorem into a deprecated alias for the more general one. --- Mathlib.lean | 8 +- Mathlib/Analysis/Convex/Radon.lean | 2 +- Mathlib/Dynamics/FixedPoints/Topology.lean | 2 +- .../Constructions/Polish/Basic.lean | 2 +- Mathlib/MeasureTheory/Measure/DiracProba.lean | 2 +- .../Topology/Algebra/InfiniteSum/Defs.lean | 2 +- Mathlib/Topology/Algebra/Semigroup.lean | 2 +- Mathlib/Topology/Algebra/Support.lean | 2 +- .../Topology/Category/Profinite/Basic.lean | 1 + .../Topology/Category/Profinite/Nobeling.lean | 1 + Mathlib/Topology/ClopenBox.lean | 1 + Mathlib/Topology/Connected/Separation.lean | 2 +- Mathlib/Topology/DenseEmbedding.lean | 2 +- Mathlib/Topology/ExtendFrom.lean | 2 +- Mathlib/Topology/JacobsonSpace.lean | 2 +- Mathlib/Topology/Order/OrderClosed.lean | 2 +- Mathlib/Topology/Order/Priestley.lean | 2 +- Mathlib/Topology/Order/T5.lean | 1 + Mathlib/Topology/Perfect.lean | 2 +- Mathlib/Topology/SeparatedMap.lean | 2 +- Mathlib/Topology/Separation/Basic.lean | 1518 +---------------- .../{ => Separation}/CompletelyRegular.lean | 0 .../CountableSeparatingOn.lean | 0 Mathlib/Topology/Separation/GDelta.lean | 2 +- Mathlib/Topology/Separation/Hausdorff.lean | 666 ++++++++ Mathlib/Topology/Separation/Profinite.lean | 121 ++ Mathlib/Topology/Separation/Regular.lean | 671 ++++++++ .../Topology/Separation/SeparatedNhds.lean | 160 ++ Mathlib/Topology/ShrinkingLemma.lean | 2 +- .../Topology/UniformSpace/OfCompactT2.lean | 2 +- Mathlib/Topology/UniformSpace/Separation.lean | 2 +- 31 files changed, 1653 insertions(+), 1533 deletions(-) rename Mathlib/Topology/{ => Separation}/CompletelyRegular.lean (100%) rename Mathlib/Topology/{ => Separation}/CountableSeparatingOn.lean (100%) create mode 100644 Mathlib/Topology/Separation/Hausdorff.lean create mode 100644 Mathlib/Topology/Separation/Profinite.lean create mode 100644 Mathlib/Topology/Separation/Regular.lean create mode 100644 Mathlib/Topology/Separation/SeparatedNhds.lean diff --git a/Mathlib.lean b/Mathlib.lean index a5829df84a252..56bf6989e9a40 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -5138,7 +5138,6 @@ import Mathlib.Topology.Compactness.LocallyCompact import Mathlib.Topology.Compactness.Paracompact import Mathlib.Topology.Compactness.PseudometrizableLindelof import Mathlib.Topology.Compactness.SigmaCompact -import Mathlib.Topology.CompletelyRegular import Mathlib.Topology.Connected.Basic import Mathlib.Topology.Connected.Clopen import Mathlib.Topology.Connected.LocallyConnected @@ -5173,7 +5172,6 @@ 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 import Mathlib.Topology.Defs.Basic import Mathlib.Topology.Defs.Filter @@ -5352,8 +5350,14 @@ import Mathlib.Topology.RestrictGen import Mathlib.Topology.Semicontinuous import Mathlib.Topology.SeparatedMap import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.Separation.CompletelyRegular +import Mathlib.Topology.Separation.CountableSeparatingOn import Mathlib.Topology.Separation.GDelta +import Mathlib.Topology.Separation.Hausdorff import Mathlib.Topology.Separation.NotNormal +import Mathlib.Topology.Separation.Profinite +import Mathlib.Topology.Separation.Regular +import Mathlib.Topology.Separation.SeparatedNhds import Mathlib.Topology.Sequences import Mathlib.Topology.Sets.Closeds import Mathlib.Topology.Sets.Compacts diff --git a/Mathlib/Analysis/Convex/Radon.lean b/Mathlib/Analysis/Convex/Radon.lean index ae64b4a1822a7..ac0312878f75b 100644 --- a/Mathlib/Analysis/Convex/Radon.lean +++ b/Mathlib/Analysis/Convex/Radon.lean @@ -6,7 +6,7 @@ Authors: Vasily Nesterov import Mathlib.Analysis.Convex.Combination import Mathlib.Data.Set.Card import Mathlib.LinearAlgebra.AffineSpace.FiniteDimensional -import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.Separation.Hausdorff /-! # Radon's theorem on convex sets diff --git a/Mathlib/Dynamics/FixedPoints/Topology.lean b/Mathlib/Dynamics/FixedPoints/Topology.lean index c657d45d24997..9c65e15446242 100644 --- a/Mathlib/Dynamics/FixedPoints/Topology.lean +++ b/Mathlib/Dynamics/FixedPoints/Topology.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Johannes Hölzl -/ import Mathlib.Dynamics.FixedPoints.Basic -import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.Separation.Hausdorff /-! # Topological properties of fixed points diff --git a/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean b/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean index cdd41441ac246..ff610fee246c5 100644 --- a/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Felix Weilacher -/ import Mathlib.MeasureTheory.Constructions.BorelSpace.Metric -import Mathlib.Topology.CountableSeparatingOn import Mathlib.Topology.MetricSpace.Perfect +import Mathlib.Topology.Separation.CountableSeparatingOn /-! # The Borel sigma-algebra on Polish spaces diff --git a/Mathlib/MeasureTheory/Measure/DiracProba.lean b/Mathlib/MeasureTheory/Measure/DiracProba.lean index 9a4f81e244015..a854d03fabb3f 100644 --- a/Mathlib/MeasureTheory/Measure/DiracProba.lean +++ b/Mathlib/MeasureTheory/Measure/DiracProba.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Kalle Kytölä. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kalle Kytölä -/ -import Mathlib.Topology.CompletelyRegular +import Mathlib.Topology.Separation.CompletelyRegular import Mathlib.MeasureTheory.Measure.ProbabilityMeasure /-! diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean b/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean index 44ab58cf88b01..db6ba150068fa 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean @@ -5,7 +5,7 @@ Authors: Johannes Hölzl -/ import Mathlib.Algebra.BigOperators.Finprod import Mathlib.Order.Filter.AtTopBot.BigOperators -import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.Separation.Hausdorff /-! # Infinite sum and product over a topological monoid diff --git a/Mathlib/Topology/Algebra/Semigroup.lean b/Mathlib/Topology/Algebra/Semigroup.lean index 327643f2aa4c3..36ed3934362d7 100644 --- a/Mathlib/Topology/Algebra/Semigroup.lean +++ b/Mathlib/Topology/Algebra/Semigroup.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: David Wärn -/ import Mathlib.Algebra.Group.Defs -import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.Separation.Regular /-! # Idempotents in topological semigroups diff --git a/Mathlib/Topology/Algebra/Support.lean b/Mathlib/Topology/Algebra/Support.lean index e11dafdd2bedd..0b2ba5734b60b 100644 --- a/Mathlib/Topology/Algebra/Support.lean +++ b/Mathlib/Topology/Algebra/Support.lean @@ -6,7 +6,7 @@ Authors: Floris van Doorn, Patrick Massot import Mathlib.Algebra.GroupWithZero.Indicator import Mathlib.Algebra.Order.Group.Unbundled.Abs import Mathlib.Algebra.Module.Basic -import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.Separation.Hausdorff /-! # The topological support of a function diff --git a/Mathlib/Topology/Category/Profinite/Basic.lean b/Mathlib/Topology/Category/Profinite/Basic.lean index e15cf996568bd..14be0aa96b9a9 100644 --- a/Mathlib/Topology/Category/Profinite/Basic.lean +++ b/Mathlib/Topology/Category/Profinite/Basic.lean @@ -6,6 +6,7 @@ Authors: Kevin Buzzard, Calle Sönne, Dagur Asgeirsson import Mathlib.CategoryTheory.FintypeCat import Mathlib.Topology.Category.CompHaus.Basic import Mathlib.Topology.LocallyConstant.Basic +import Mathlib.Topology.Separation.Profinite /-! # The category of Profinite Types diff --git a/Mathlib/Topology/Category/Profinite/Nobeling.lean b/Mathlib/Topology/Category/Profinite/Nobeling.lean index c244872abc715..494a50eb97186 100644 --- a/Mathlib/Topology/Category/Profinite/Nobeling.lean +++ b/Mathlib/Topology/Category/Profinite/Nobeling.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.Category.ModuleCat.Free import Mathlib.Topology.Category.Profinite.CofilteredLimit import Mathlib.Topology.Category.Profinite.Product import Mathlib.Topology.LocallyConstant.Algebra +import Mathlib.Topology.Separation.Profinite import Mathlib.Data.Bool.Basic /-! diff --git a/Mathlib/Topology/ClopenBox.lean b/Mathlib/Topology/ClopenBox.lean index a5b66dd989ec4..07d40b673ec08 100644 --- a/Mathlib/Topology/ClopenBox.lean +++ b/Mathlib/Topology/ClopenBox.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Dagur Asgeirsson -/ import Mathlib.Topology.CompactOpen +import Mathlib.Topology.Separation.Profinite import Mathlib.Topology.Sets.Closeds /-! diff --git a/Mathlib/Topology/Connected/Separation.lean b/Mathlib/Topology/Connected/Separation.lean index f5e3e067c4839..9cb603e6e6d64 100644 --- a/Mathlib/Topology/Connected/Separation.lean +++ b/Mathlib/Topology/Connected/Separation.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Dagur Asgeirsson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Dagur Asgeirsson -/ -import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.Separation.Hausdorff /-! # Separation and (dis)connectedness properties of topological spaces. diff --git a/Mathlib/Topology/DenseEmbedding.lean b/Mathlib/Topology/DenseEmbedding.lean index d615811b191db..ac53d77bf26a1 100644 --- a/Mathlib/Topology/DenseEmbedding.lean +++ b/Mathlib/Topology/DenseEmbedding.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.Topology.Bases -import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.Separation.Regular /-! # Dense embeddings diff --git a/Mathlib/Topology/ExtendFrom.lean b/Mathlib/Topology/ExtendFrom.lean index 015142b528778..649538b26ec69 100644 --- a/Mathlib/Topology/ExtendFrom.lean +++ b/Mathlib/Topology/ExtendFrom.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Anatole Dedecker. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Anatole Dedecker -/ -import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.Separation.Regular /-! # Extending a function from a subset diff --git a/Mathlib/Topology/JacobsonSpace.lean b/Mathlib/Topology/JacobsonSpace.lean index 10bdb54e1e515..2e03eba48dbca 100644 --- a/Mathlib/Topology/JacobsonSpace.lean +++ b/Mathlib/Topology/JacobsonSpace.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ import Mathlib.Topology.LocalAtTarget -import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.Separation.Regular import Mathlib.Tactic.StacksAttribute /-! diff --git a/Mathlib/Topology/Order/OrderClosed.lean b/Mathlib/Topology/Order/OrderClosed.lean index 46e8626f54295..226ff0babfa1c 100644 --- a/Mathlib/Topology/Order/OrderClosed.lean +++ b/Mathlib/Topology/Order/OrderClosed.lean @@ -3,7 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Yury Kudryashov -/ -import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.Separation.Hausdorff /-! # Order-closed topologies diff --git a/Mathlib/Topology/Order/Priestley.lean b/Mathlib/Topology/Order/Priestley.lean index bdad508392e4a..ea4770282d3d5 100644 --- a/Mathlib/Topology/Order/Priestley.lean +++ b/Mathlib/Topology/Order/Priestley.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Order.UpperLower.Basic -import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.Separation.Hausdorff /-! # Priestley spaces diff --git a/Mathlib/Topology/Order/T5.lean b/Mathlib/Topology/Order/T5.lean index 597085db3f109..fadf00fc219bc 100644 --- a/Mathlib/Topology/Order/T5.lean +++ b/Mathlib/Topology/Order/T5.lean @@ -5,6 +5,7 @@ Authors: Yury Kudryashov -/ import Mathlib.Order.Interval.Set.OrdConnectedComponent import Mathlib.Topology.Order.Basic +import Mathlib.Topology.Separation.Regular /-! # Linear order is a completely normal Hausdorff topological space diff --git a/Mathlib/Topology/Perfect.lean b/Mathlib/Topology/Perfect.lean index 88e7d579f3dd7..5c03ddd913ef5 100644 --- a/Mathlib/Topology/Perfect.lean +++ b/Mathlib/Topology/Perfect.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Felix Weilacher -/ -import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.Separation.Regular /-! # Perfect Sets diff --git a/Mathlib/Topology/SeparatedMap.lean b/Mathlib/Topology/SeparatedMap.lean index 4d37a595ce392..61b7da0770428 100644 --- a/Mathlib/Topology/SeparatedMap.lean +++ b/Mathlib/Topology/SeparatedMap.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Junyan Xu -/ import Mathlib.Topology.Connected.Basic -import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.Separation.Hausdorff /-! # Separated maps and locally injective maps out of a topological space. diff --git a/Mathlib/Topology/Separation/Basic.lean b/Mathlib/Topology/Separation/Basic.lean index 827ce1e31923c..9f6cb623f903b 100644 --- a/Mathlib/Topology/Separation/Basic.lean +++ b/Mathlib/Topology/Separation/Basic.lean @@ -4,16 +4,17 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ import Mathlib.Algebra.Group.Support -import Mathlib.Topology.Compactness.Lindelof -import Mathlib.Topology.Compactness.SigmaCompact import Mathlib.Topology.Connected.TotallyDisconnected import Mathlib.Topology.Inseparable +import Mathlib.Topology.Separation.SeparatedNhds +import Mathlib.Topology.Compactness.LocallyCompact /-! -# Separation properties of topological spaces. +# Separation properties of topological spaces -This file defines the predicate `SeparatedNhds`, and common separation axioms -(under the Kolmogorov classification). +This file defines some of the weaker separation axioms (under the Kolmogorov classification), +notably T₀, R₀, T₁ and R₁ spaces. For T₂ (Hausdorff) spaces and other stronger +conditions, see the file `Topology/Separation/Hausdorff.lean`. ## Main definitions @@ -31,28 +32,6 @@ This file defines the predicate `SeparatedNhds`, and common separation axioms T₁ implies T₀ and R₀. * `R1Space`: An R₁/preregular space is a space where any two topologically distinguishable points have disjoint neighbourhoods. R₁ implies R₀. -* `T2Space`: A T₂/Hausdorff space is a space where, for every two points `x ≠ y`, - there is two disjoint open sets, one containing `x`, and the other `y`. T₂ implies T₁ and R₁. -* `T25Space`: A T₂.₅/Urysohn space is a space where, for every two points `x ≠ y`, - there is two open sets, one containing `x`, and the other `y`, whose closures are disjoint. - T₂.₅ implies T₂. -* `RegularSpace`: A regular space is one where, given any closed `C` and `x ∉ C`, - there are disjoint open sets containing `x` and `C` respectively. Such a space is not necessarily - Hausdorff. -* `T3Space`: A T₃ space is a regular T₀ space. T₃ implies T₂.₅. -* `NormalSpace`: A normal space, is one where given two disjoint closed sets, - we can find two open sets that separate them. Such a space is not necessarily Hausdorff, even if - it is T₀. -* `T4Space`: A T₄ space is a normal T₁ space. T₄ implies T₃. -* `CompletelyNormalSpace`: A completely normal space is one in which for any two sets `s`, `t` - such that if both `closure s` is disjoint with `t`, and `s` is disjoint with `closure t`, - then there exist disjoint neighbourhoods of `s` and `t`. `Embedding.completelyNormalSpace` allows - us to conclude that this is equivalent to all subspaces being normal. Such a space is not - necessarily Hausdorff or regular, even if it is T₀. -* `T5Space`: A T₅ space is a completely normal T₁ space. T₅ implies T₄. - -See `Mathlib.Topology.Separation.GDelta` for the definitions of `PerfectlyNormalSpace` and -`T6Space`. Note that `mathlib` adopts the modern convention that `m ≤ n` if and only if `T_m → T_n`, but occasionally the literature swaps definitions for e.g. T₃ and regular. @@ -71,55 +50,12 @@ occasionally the literature swaps definitions for e.g. T₃ and regular. * `isClosedMap_const`: The constant map is a closed map. * `Finite.instDiscreteTopology`: A finite T₁ space must have the discrete topology. -### T₂ spaces - -* `t2_iff_nhds`: A space is T₂ iff the neighbourhoods of distinct points generate the bottom filter. -* `t2_iff_isClosed_diagonal`: A space is T₂ iff the `diagonal` of `X` (that is, the set of all - points of the form `(a, a) : X × X`) is closed under the product topology. -* `separatedNhds_of_finset_finset`: Any two disjoint finsets are `SeparatedNhds`. -* Most topological constructions preserve Hausdorffness; - these results are part of the typeclass inference system (e.g. `Topology.IsEmbedding.t2Space`) -* `Set.EqOn.closure`: If two functions are equal on some set `s`, they are equal on its closure. -* `IsCompact.isClosed`: All compact sets are closed. -* `WeaklyLocallyCompactSpace.locallyCompactSpace`: If a topological space is both - weakly locally compact (i.e., each point has a compact neighbourhood) - and is T₂, then it is locally compact. -* `totallySeparatedSpace_of_t1_of_basis_clopen`: If `X` has a clopen basis, then - it is a `TotallySeparatedSpace`. -* `loc_compact_t2_tot_disc_iff_tot_sep`: A locally compact T₂ space is totally disconnected iff - it is totally separated. -* `t2Quotient`: the largest T2 quotient of a given topological space. - -If the space is also compact: - -* `normalOfCompactT2`: A compact T₂ space is a `NormalSpace`. -* `connectedComponent_eq_iInter_isClopen`: The connected component of a point - is the intersection of all its clopen neighbourhoods. -* `compact_t2_tot_disc_iff_tot_sep`: Being a `TotallyDisconnectedSpace` - is equivalent to being a `TotallySeparatedSpace`. -* `ConnectedComponents.t2`: `ConnectedComponents X` is T₂ for `X` T₂ and compact. - -### Regular spaces - -If the space is also Lindelöf: - -* `NormalSpace.of_regularSpace_lindelofSpace`: every regular Lindelöf space is normal. - -### T₃ spaces - -* `disjoint_nested_nhds`: Given two points `x ≠ y`, we can find neighbourhoods `x ∈ V₁ ⊆ U₁` and - `y ∈ V₂ ⊆ U₂`, with the `Vₖ` closed and the `Uₖ` open, such that the `Uₖ` are disjoint. - ## References * -* * [Willard's *General Topology*][zbMATH02107988] - -/ -assert_not_exists UniformSpace - open Function Set Filter Topology TopologicalSpace universe u v @@ -128,131 +64,6 @@ variable {X : Type*} {Y : Type*} [TopologicalSpace X] section Separation -/-- -`SeparatedNhds` is a predicate on pairs of sub`Set`s of a topological space. It holds if the two -sub`Set`s are contained in disjoint open sets. --/ -def SeparatedNhds : Set X → Set X → Prop := fun s t : Set X => - ∃ U V : Set X, IsOpen U ∧ IsOpen V ∧ s ⊆ U ∧ t ⊆ V ∧ Disjoint U V - -theorem separatedNhds_iff_disjoint {s t : Set X} : SeparatedNhds s t ↔ Disjoint (𝓝ˢ s) (𝓝ˢ t) := by - simp only [(hasBasis_nhdsSet s).disjoint_iff (hasBasis_nhdsSet t), SeparatedNhds, exists_prop, ← - exists_and_left, and_assoc, and_comm, and_left_comm] - -alias ⟨SeparatedNhds.disjoint_nhdsSet, _⟩ := separatedNhds_iff_disjoint - -/-- `HasSeparatingCover`s can be useful witnesses for `SeparatedNhds`. -/ -def HasSeparatingCover : Set X → Set X → Prop := fun s t ↦ - ∃ u : ℕ → Set X, s ⊆ ⋃ n, u n ∧ ∀ n, IsOpen (u n) ∧ Disjoint (closure (u n)) t - -/-- Used to prove that a regular topological space with Lindelöf topology is a normal space, -and a perfectly normal space is a completely normal space. -/ -theorem hasSeparatingCovers_iff_separatedNhds {s t : Set X} : - HasSeparatingCover s t ∧ HasSeparatingCover t s ↔ SeparatedNhds s t := by - constructor - · rintro ⟨⟨u, u_cov, u_props⟩, ⟨v, v_cov, v_props⟩⟩ - have open_lemma : ∀ (u₀ a : ℕ → Set X), (∀ n, IsOpen (u₀ n)) → - IsOpen (⋃ n, u₀ n \ closure (a n)) := fun _ _ u₀i_open ↦ - isOpen_iUnion fun i ↦ (u₀i_open i).sdiff isClosed_closure - have cover_lemma : ∀ (h₀ : Set X) (u₀ v₀ : ℕ → Set X), - (h₀ ⊆ ⋃ n, u₀ n) → (∀ n, Disjoint (closure (v₀ n)) h₀) → - (h₀ ⊆ ⋃ n, u₀ n \ closure (⋃ m ≤ n, v₀ m)) := - fun h₀ u₀ v₀ h₀_cov dis x xinh ↦ by - rcases h₀_cov xinh with ⟨un , ⟨n, rfl⟩ , xinun⟩ - simp only [mem_iUnion] - refine ⟨n, xinun, ?_⟩ - simp_all only [closure_iUnion₂_le_nat, disjoint_right, mem_setOf_eq, mem_iUnion, - exists_false, exists_const, not_false_eq_true] - refine - ⟨⋃ n : ℕ, u n \ (closure (⋃ m ≤ n, v m)), - ⋃ n : ℕ, v n \ (closure (⋃ m ≤ n, u m)), - open_lemma u (fun n ↦ ⋃ m ≤ n, v m) (fun n ↦ (u_props n).1), - open_lemma v (fun n ↦ ⋃ m ≤ n, u m) (fun n ↦ (v_props n).1), - cover_lemma s u v u_cov (fun n ↦ (v_props n).2), - cover_lemma t v u v_cov (fun n ↦ (u_props n).2), - ?_⟩ - rw [Set.disjoint_left] - rintro x ⟨un, ⟨n, rfl⟩, xinun⟩ - suffices ∀ (m : ℕ), x ∈ v m → x ∈ closure (⋃ m' ∈ {m' | m' ≤ m}, u m') by simpa - intro m xinvm - have n_le_m : n ≤ m := by - by_contra m_gt_n - exact xinun.2 (subset_closure (mem_biUnion (le_of_lt (not_le.mp m_gt_n)) xinvm)) - exact subset_closure (mem_biUnion n_le_m xinun.1) - · rintro ⟨U, V, U_open, V_open, h_sub_U, k_sub_V, UV_dis⟩ - exact - ⟨⟨fun _ ↦ U, - h_sub_U.trans (iUnion_const U).symm.subset, - fun _ ↦ - ⟨U_open, disjoint_of_subset (fun ⦃a⦄ a ↦ a) k_sub_V (UV_dis.closure_left V_open)⟩⟩, - ⟨fun _ ↦ V, - k_sub_V.trans (iUnion_const V).symm.subset, - fun _ ↦ - ⟨V_open, disjoint_of_subset (fun ⦃a⦄ a ↦ a) h_sub_U (UV_dis.closure_right U_open).symm⟩⟩⟩ - -theorem Set.hasSeparatingCover_empty_left (s : Set X) : HasSeparatingCover ∅ s := - ⟨fun _ ↦ ∅, empty_subset (⋃ _, ∅), - fun _ ↦ ⟨isOpen_empty, by simp only [closure_empty, empty_disjoint]⟩⟩ - -theorem Set.hasSeparatingCover_empty_right (s : Set X) : HasSeparatingCover s ∅ := - ⟨fun _ ↦ univ, (subset_univ s).trans univ.iUnion_const.symm.subset, - fun _ ↦ ⟨isOpen_univ, by apply disjoint_empty⟩⟩ - -theorem HasSeparatingCover.mono {s₁ s₂ t₁ t₂ : Set X} (sc_st : HasSeparatingCover s₂ t₂) - (s_sub : s₁ ⊆ s₂) (t_sub : t₁ ⊆ t₂) : HasSeparatingCover s₁ t₁ := by - obtain ⟨u, u_cov, u_props⟩ := sc_st - exact - ⟨u, - s_sub.trans u_cov, - fun n ↦ - ⟨(u_props n).1, - disjoint_of_subset (fun ⦃_⦄ a ↦ a) t_sub (u_props n).2⟩⟩ - -namespace SeparatedNhds - -variable {s s₁ s₂ t t₁ t₂ u : Set X} - -@[symm] -theorem symm : SeparatedNhds s t → SeparatedNhds t s := fun ⟨U, V, oU, oV, aU, bV, UV⟩ => - ⟨V, U, oV, oU, bV, aU, Disjoint.symm UV⟩ - -theorem comm (s t : Set X) : SeparatedNhds s t ↔ SeparatedNhds t s := - ⟨symm, symm⟩ - -theorem preimage [TopologicalSpace Y] {f : X → Y} {s t : Set Y} (h : SeparatedNhds s t) - (hf : Continuous f) : SeparatedNhds (f ⁻¹' s) (f ⁻¹' t) := - let ⟨U, V, oU, oV, sU, tV, UV⟩ := h - ⟨f ⁻¹' U, f ⁻¹' V, oU.preimage hf, oV.preimage hf, preimage_mono sU, preimage_mono tV, - UV.preimage f⟩ - -protected theorem disjoint (h : SeparatedNhds s t) : Disjoint s t := - let ⟨_, _, _, _, hsU, htV, hd⟩ := h; hd.mono hsU htV - -theorem disjoint_closure_left (h : SeparatedNhds s t) : Disjoint (closure s) t := - let ⟨_U, _V, _, hV, hsU, htV, hd⟩ := h - (hd.closure_left hV).mono (closure_mono hsU) htV - -theorem disjoint_closure_right (h : SeparatedNhds s t) : Disjoint s (closure t) := - h.symm.disjoint_closure_left.symm - -@[simp] theorem empty_right (s : Set X) : SeparatedNhds s ∅ := - ⟨_, _, isOpen_univ, isOpen_empty, fun a _ => mem_univ a, Subset.rfl, disjoint_empty _⟩ - -@[simp] theorem empty_left (s : Set X) : SeparatedNhds ∅ s := - (empty_right _).symm - -theorem mono (h : SeparatedNhds s₂ t₂) (hs : s₁ ⊆ s₂) (ht : t₁ ⊆ t₂) : SeparatedNhds s₁ t₁ := - let ⟨U, V, hU, hV, hsU, htV, hd⟩ := h - ⟨U, V, hU, hV, hs.trans hsU, ht.trans htV, hd⟩ - -theorem union_left : SeparatedNhds s u → SeparatedNhds t u → SeparatedNhds (s ∪ t) u := by - simpa only [separatedNhds_iff_disjoint, nhdsSet_union, disjoint_sup_left] using And.intro - -theorem union_right (ht : SeparatedNhds s t) (hu : SeparatedNhds s u) : SeparatedNhds s (t ∪ u) := - (ht.symm.union_left hu.symm).symm - -end SeparatedNhds - /-- A T₀ space, also known as a Kolmogorov space, is a topological space such that for every pair `x ≠ y`, there is an open set containing one but not the other. We formulate the definition in terms of the `Inseparable` relation. -/ @@ -1302,1321 +1113,4 @@ theorem exists_isOpen_mem_isCompact_closure (x : X) : end R1Space -/-- A T₂ space, also known as a Hausdorff space, is one in which for every - `x ≠ y` there exists disjoint open sets around `x` and `y`. This is - the most widely used of the separation axioms. -/ -@[mk_iff] -class T2Space (X : Type u) [TopologicalSpace X] : Prop where - /-- Every two points in a Hausdorff space admit disjoint open neighbourhoods. -/ - t2 : Pairwise fun x y => ∃ u v : Set X, IsOpen u ∧ IsOpen v ∧ x ∈ u ∧ y ∈ v ∧ Disjoint u v - -/-- Two different points can be separated by open sets. -/ -theorem t2_separation [T2Space X] {x y : X} (h : x ≠ y) : - ∃ u v : Set X, IsOpen u ∧ IsOpen v ∧ x ∈ u ∧ y ∈ v ∧ Disjoint u v := - T2Space.t2 h - --- todo: use this as a definition? -theorem t2Space_iff_disjoint_nhds : T2Space X ↔ Pairwise fun x y : X => Disjoint (𝓝 x) (𝓝 y) := by - refine (t2Space_iff X).trans (forall₃_congr fun x y _ => ?_) - simp only [(nhds_basis_opens x).disjoint_iff (nhds_basis_opens y), exists_prop, ← exists_and_left, - and_assoc, and_comm, and_left_comm] - -@[simp] -theorem disjoint_nhds_nhds [T2Space X] {x y : X} : Disjoint (𝓝 x) (𝓝 y) ↔ x ≠ y := - ⟨fun hd he => by simp [he, nhds_neBot.ne] at hd, (t2Space_iff_disjoint_nhds.mp ‹_› ·)⟩ - -theorem pairwise_disjoint_nhds [T2Space X] : Pairwise (Disjoint on (𝓝 : X → Filter X)) := fun _ _ => - disjoint_nhds_nhds.2 - -protected theorem Set.pairwiseDisjoint_nhds [T2Space X] (s : Set X) : s.PairwiseDisjoint 𝓝 := - pairwise_disjoint_nhds.set_pairwise s - -/-- Points of a finite set can be separated by open sets from each other. -/ -theorem Set.Finite.t2_separation [T2Space X] {s : Set X} (hs : s.Finite) : - ∃ U : X → Set X, (∀ x, x ∈ U x ∧ IsOpen (U x)) ∧ s.PairwiseDisjoint U := - s.pairwiseDisjoint_nhds.exists_mem_filter_basis hs nhds_basis_opens - --- see Note [lower instance priority] -instance (priority := 100) T2Space.t1Space [T2Space X] : T1Space X := - t1Space_iff_disjoint_pure_nhds.mpr fun _ _ hne => - (disjoint_nhds_nhds.2 hne).mono_left <| pure_le_nhds _ - --- see Note [lower instance priority] -instance (priority := 100) T2Space.r1Space [T2Space X] : R1Space X := - ⟨fun x y ↦ (eq_or_ne x y).imp specializes_of_eq disjoint_nhds_nhds.2⟩ - -theorem SeparationQuotient.t2Space_iff : T2Space (SeparationQuotient X) ↔ R1Space X := by - simp only [t2Space_iff_disjoint_nhds, Pairwise, surjective_mk.forall₂, ne_eq, mk_eq_mk, - r1Space_iff_inseparable_or_disjoint_nhds, ← disjoint_comap_iff surjective_mk, comap_mk_nhds_mk, - ← or_iff_not_imp_left] - -instance SeparationQuotient.t2Space [R1Space X] : T2Space (SeparationQuotient X) := - t2Space_iff.2 ‹_› - -instance (priority := 80) [R1Space X] [T0Space X] : T2Space X := - t2Space_iff_disjoint_nhds.2 fun _x _y hne ↦ disjoint_nhds_nhds_iff_not_inseparable.2 fun hxy ↦ - hne hxy.eq - -theorem R1Space.t2Space_iff_t0Space [R1Space X] : T2Space X ↔ T0Space X := by - constructor <;> intro <;> infer_instance - -/-- A space is T₂ iff the neighbourhoods of distinct points generate the bottom filter. -/ -theorem t2_iff_nhds : T2Space X ↔ ∀ {x y : X}, NeBot (𝓝 x ⊓ 𝓝 y) → x = y := by - simp only [t2Space_iff_disjoint_nhds, disjoint_iff, neBot_iff, Ne, not_imp_comm, Pairwise] - -theorem eq_of_nhds_neBot [T2Space X] {x y : X} (h : NeBot (𝓝 x ⊓ 𝓝 y)) : x = y := - t2_iff_nhds.mp ‹_› h - -theorem t2Space_iff_nhds : - T2Space X ↔ Pairwise fun x y : X => ∃ U ∈ 𝓝 x, ∃ V ∈ 𝓝 y, Disjoint U V := by - simp only [t2Space_iff_disjoint_nhds, Filter.disjoint_iff, Pairwise] - -theorem t2_separation_nhds [T2Space X] {x y : X} (h : x ≠ y) : - ∃ u v, u ∈ 𝓝 x ∧ v ∈ 𝓝 y ∧ Disjoint u v := - let ⟨u, v, open_u, open_v, x_in, y_in, huv⟩ := t2_separation h - ⟨u, v, open_u.mem_nhds x_in, open_v.mem_nhds y_in, huv⟩ - -theorem t2_separation_compact_nhds [LocallyCompactSpace X] [T2Space X] {x y : X} (h : x ≠ y) : - ∃ u v, u ∈ 𝓝 x ∧ v ∈ 𝓝 y ∧ IsCompact u ∧ IsCompact v ∧ Disjoint u v := by - simpa only [exists_prop, ← exists_and_left, and_comm, and_assoc, and_left_comm] using - ((compact_basis_nhds x).disjoint_iff (compact_basis_nhds y)).1 (disjoint_nhds_nhds.2 h) - -theorem t2_iff_ultrafilter : - T2Space X ↔ ∀ {x y : X} (f : Ultrafilter X), ↑f ≤ 𝓝 x → ↑f ≤ 𝓝 y → x = y := - t2_iff_nhds.trans <| by simp only [← exists_ultrafilter_iff, and_imp, le_inf_iff, exists_imp] - -theorem t2_iff_isClosed_diagonal : T2Space X ↔ IsClosed (diagonal X) := by - simp only [t2Space_iff_disjoint_nhds, ← isOpen_compl_iff, isOpen_iff_mem_nhds, Prod.forall, - nhds_prod_eq, compl_diagonal_mem_prod, mem_compl_iff, mem_diagonal_iff, Pairwise] - -theorem isClosed_diagonal [T2Space X] : IsClosed (diagonal X) := - t2_iff_isClosed_diagonal.mp ‹_› - --- Porting note: 2 lemmas moved below - -theorem tendsto_nhds_unique [T2Space X] {f : Y → X} {l : Filter Y} {a b : X} [NeBot l] - (ha : Tendsto f l (𝓝 a)) (hb : Tendsto f l (𝓝 b)) : a = b := - (tendsto_nhds_unique_inseparable ha hb).eq - -theorem tendsto_nhds_unique' [T2Space X] {f : Y → X} {l : Filter Y} {a b : X} (_ : NeBot l) - (ha : Tendsto f l (𝓝 a)) (hb : Tendsto f l (𝓝 b)) : a = b := - tendsto_nhds_unique ha hb - -theorem tendsto_nhds_unique_of_eventuallyEq [T2Space X] {f g : Y → X} {l : Filter Y} {a b : X} - [NeBot l] (ha : Tendsto f l (𝓝 a)) (hb : Tendsto g l (𝓝 b)) (hfg : f =ᶠ[l] g) : a = b := - tendsto_nhds_unique (ha.congr' hfg) hb - -theorem tendsto_nhds_unique_of_frequently_eq [T2Space X] {f g : Y → X} {l : Filter Y} {a b : X} - (ha : Tendsto f l (𝓝 a)) (hb : Tendsto g l (𝓝 b)) (hfg : ∃ᶠ x in l, f x = g x) : a = b := - have : ∃ᶠ z : X × X in 𝓝 (a, b), z.1 = z.2 := (ha.prod_mk_nhds hb).frequently hfg - not_not.1 fun hne => this (isClosed_diagonal.isOpen_compl.mem_nhds hne) - -/-- If `s` and `t` are compact sets in a T₂ space, then the set neighborhoods filter of `s ∩ t` -is the infimum of set neighborhoods filters for `s` and `t`. - -For general sets, only the `≤` inequality holds, see `nhdsSet_inter_le`. -/ -theorem IsCompact.nhdsSet_inter_eq [T2Space X] {s t : Set X} (hs : IsCompact s) (ht : IsCompact t) : - 𝓝ˢ (s ∩ t) = 𝓝ˢ s ⊓ 𝓝ˢ t := by - refine le_antisymm (nhdsSet_inter_le _ _) ?_ - simp_rw [hs.nhdsSet_inf_eq_biSup, ht.inf_nhdsSet_eq_biSup, nhdsSet, sSup_image] - refine iSup₂_le fun x hxs ↦ iSup₂_le fun y hyt ↦ ?_ - rcases eq_or_ne x y with (rfl|hne) - · exact le_iSup₂_of_le x ⟨hxs, hyt⟩ (inf_idem _).le - · exact (disjoint_nhds_nhds.mpr hne).eq_bot ▸ bot_le - -/-- In a `T2Space X`, for a compact set `t` and a point `x` outside `t`, there are open sets `U`, -`V` that separate `t` and `x`.-/ -lemma IsCompact.separation_of_not_mem {X : Type u_1} [TopologicalSpace X] [T2Space X] {x : X} - {t : Set X} (H1 : IsCompact t) (H2 : x ∉ t) : - ∃ (U : Set X), ∃ (V : Set X), IsOpen U ∧ IsOpen V ∧ t ⊆ U ∧ x ∈ V ∧ Disjoint U V := by - simpa [SeparatedNhds] using SeparatedNhds.of_isCompact_isCompact_isClosed H1 isCompact_singleton - isClosed_singleton <| disjoint_singleton_right.mpr H2 - -/-- In a `T2Space X`, for a compact set `t` and a point `x` outside `t`, `𝓝ˢ t` and `𝓝 x` are -disjoint. -/ -lemma IsCompact.disjoint_nhdsSet_nhds {X : Type u_1} [TopologicalSpace X] [T2Space X] {x : X} - {t : Set X} (H1 : IsCompact t) (H2 : x ∉ t) : - Disjoint (𝓝ˢ t) (𝓝 x) := by - simpa using SeparatedNhds.disjoint_nhdsSet <| .of_isCompact_isCompact_isClosed H1 - isCompact_singleton isClosed_singleton <| disjoint_singleton_right.mpr H2 - -/-- If a function `f` is - -- injective on a compact set `s`; -- continuous at every point of this set; -- injective on a neighborhood of each point of this set, - -then it is injective on a neighborhood of this set. -/ -theorem Set.InjOn.exists_mem_nhdsSet {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] - [T2Space Y] {f : X → Y} {s : Set X} (inj : InjOn f s) (sc : IsCompact s) - (fc : ∀ x ∈ s, ContinuousAt f x) (loc : ∀ x ∈ s, ∃ u ∈ 𝓝 x, InjOn f u) : - ∃ t ∈ 𝓝ˢ s, InjOn f t := by - have : ∀ x ∈ s ×ˢ s, ∀ᶠ y in 𝓝 x, f y.1 = f y.2 → y.1 = y.2 := fun (x, y) ⟨hx, hy⟩ ↦ by - rcases eq_or_ne x y with rfl | hne - · rcases loc x hx with ⟨u, hu, hf⟩ - exact Filter.mem_of_superset (prod_mem_nhds hu hu) <| forall_prod_set.2 hf - · suffices ∀ᶠ z in 𝓝 (x, y), f z.1 ≠ f z.2 from this.mono fun _ hne h ↦ absurd h hne - refine (fc x hx).prodMap' (fc y hy) <| isClosed_diagonal.isOpen_compl.mem_nhds ?_ - exact inj.ne hx hy hne - rw [← eventually_nhdsSet_iff_forall, sc.nhdsSet_prod_eq sc] at this - exact eventually_prod_self_iff.1 this - -/-- If a function `f` is - -- injective on a compact set `s`; -- continuous at every point of this set; -- injective on a neighborhood of each point of this set, - -then it is injective on an open neighborhood of this set. -/ -theorem Set.InjOn.exists_isOpen_superset {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] - [T2Space Y] {f : X → Y} {s : Set X} (inj : InjOn f s) (sc : IsCompact s) - (fc : ∀ x ∈ s, ContinuousAt f x) (loc : ∀ x ∈ s, ∃ u ∈ 𝓝 x, InjOn f u) : - ∃ t, IsOpen t ∧ s ⊆ t ∧ InjOn f t := - let ⟨_t, hst, ht⟩ := inj.exists_mem_nhdsSet sc fc loc - let ⟨u, huo, hsu, hut⟩ := mem_nhdsSet_iff_exists.1 hst - ⟨u, huo, hsu, ht.mono hut⟩ - -section limUnder - -variable [T2Space X] {f : Filter X} - -/-! -### Properties of `lim` and `limUnder` - -In this section we use explicit `Nonempty X` instances for `lim` and `limUnder`. This way the lemmas -are useful without a `Nonempty X` instance. --/ - - -theorem lim_eq {x : X} [NeBot f] (h : f ≤ 𝓝 x) : @lim _ _ ⟨x⟩ f = x := - tendsto_nhds_unique (le_nhds_lim ⟨x, h⟩) h - -theorem lim_eq_iff [NeBot f] (h : ∃ x : X, f ≤ 𝓝 x) {x} : @lim _ _ ⟨x⟩ f = x ↔ f ≤ 𝓝 x := - ⟨fun c => c ▸ le_nhds_lim h, lim_eq⟩ - -theorem Ultrafilter.lim_eq_iff_le_nhds [CompactSpace X] {x : X} {F : Ultrafilter X} : - F.lim = x ↔ ↑F ≤ 𝓝 x := - ⟨fun h => h ▸ F.le_nhds_lim, lim_eq⟩ - -theorem isOpen_iff_ultrafilter' [CompactSpace X] (U : Set X) : - IsOpen U ↔ ∀ F : Ultrafilter X, F.lim ∈ U → U ∈ F.1 := by - rw [isOpen_iff_ultrafilter] - refine ⟨fun h F hF => h F.lim hF F F.le_nhds_lim, ?_⟩ - intro cond x hx f h - rw [← Ultrafilter.lim_eq_iff_le_nhds.2 h] at hx - exact cond _ hx - -theorem Filter.Tendsto.limUnder_eq {x : X} {f : Filter Y} [NeBot f] {g : Y → X} - (h : Tendsto g f (𝓝 x)) : @limUnder _ _ _ ⟨x⟩ f g = x := - lim_eq h - -theorem Filter.limUnder_eq_iff {f : Filter Y} [NeBot f] {g : Y → X} (h : ∃ x, Tendsto g f (𝓝 x)) - {x} : @limUnder _ _ _ ⟨x⟩ f g = x ↔ Tendsto g f (𝓝 x) := - ⟨fun c => c ▸ tendsto_nhds_limUnder h, Filter.Tendsto.limUnder_eq⟩ - -theorem Continuous.limUnder_eq [TopologicalSpace Y] {f : Y → X} (h : Continuous f) (y : Y) : - @limUnder _ _ _ ⟨f y⟩ (𝓝 y) f = f y := - (h.tendsto y).limUnder_eq - -@[simp] -theorem lim_nhds (x : X) : @lim _ _ ⟨x⟩ (𝓝 x) = x := - lim_eq le_rfl - -@[simp] -theorem limUnder_nhds_id (x : X) : @limUnder _ _ _ ⟨x⟩ (𝓝 x) id = x := - lim_nhds x - -@[simp] -theorem lim_nhdsWithin {x : X} {s : Set X} (h : x ∈ closure s) : @lim _ _ ⟨x⟩ (𝓝[s] x) = x := - haveI : NeBot (𝓝[s] x) := mem_closure_iff_clusterPt.1 h - lim_eq inf_le_left - -@[simp] -theorem limUnder_nhdsWithin_id {x : X} {s : Set X} (h : x ∈ closure s) : - @limUnder _ _ _ ⟨x⟩ (𝓝[s] x) id = x := - lim_nhdsWithin h - -end limUnder - -/-! -### `T2Space` constructions - -We use two lemmas to prove that various standard constructions generate Hausdorff spaces from -Hausdorff spaces: - -* `separated_by_continuous` says that two points `x y : X` can be separated by open neighborhoods - provided that there exists a continuous map `f : X → Y` with a Hausdorff codomain such that - `f x ≠ f y`. We use this lemma to prove that topological spaces defined using `induced` are - Hausdorff spaces. - -* `separated_by_isOpenEmbedding` says that for an open embedding `f : X → Y` of a Hausdorff space - `X`, the images of two distinct points `x y : X`, `x ≠ y` can be separated by open neighborhoods. - We use this lemma to prove that topological spaces defined using `coinduced` are Hausdorff spaces. --/ - --- see Note [lower instance priority] -instance (priority := 100) DiscreteTopology.toT2Space - [DiscreteTopology X] : T2Space X := - ⟨fun x y h => ⟨{x}, {y}, isOpen_discrete _, isOpen_discrete _, rfl, rfl, disjoint_singleton.2 h⟩⟩ - -theorem separated_by_continuous [TopologicalSpace Y] [T2Space Y] - {f : X → Y} (hf : Continuous f) {x y : X} (h : f x ≠ f y) : - ∃ u v : Set X, IsOpen u ∧ IsOpen v ∧ x ∈ u ∧ y ∈ v ∧ Disjoint u v := - let ⟨u, v, uo, vo, xu, yv, uv⟩ := t2_separation h - ⟨f ⁻¹' u, f ⁻¹' v, uo.preimage hf, vo.preimage hf, xu, yv, uv.preimage _⟩ - -theorem separated_by_isOpenEmbedding [TopologicalSpace Y] [T2Space X] - {f : X → Y} (hf : IsOpenEmbedding f) {x y : X} (h : x ≠ y) : - ∃ u v : Set Y, IsOpen u ∧ IsOpen v ∧ f x ∈ u ∧ f y ∈ v ∧ Disjoint u v := - let ⟨u, v, uo, vo, xu, yv, uv⟩ := t2_separation h - ⟨f '' u, f '' v, hf.isOpenMap _ uo, hf.isOpenMap _ vo, mem_image_of_mem _ xu, - mem_image_of_mem _ yv, disjoint_image_of_injective hf.injective uv⟩ - -@[deprecated (since := "2024-10-18")] -alias separated_by_openEmbedding := separated_by_isOpenEmbedding - -instance {p : X → Prop} [T2Space X] : T2Space (Subtype p) := inferInstance - -instance Prod.t2Space [T2Space X] [TopologicalSpace Y] [T2Space Y] : T2Space (X × Y) := - inferInstance - -/-- If the codomain of an injective continuous function is a Hausdorff space, then so is its -domain. -/ -theorem T2Space.of_injective_continuous [TopologicalSpace Y] [T2Space Y] {f : X → Y} - (hinj : Injective f) (hc : Continuous f) : T2Space X := - ⟨fun _ _ h => separated_by_continuous hc (hinj.ne h)⟩ - -/-- If the codomain of a topological embedding is a Hausdorff space, then so is its domain. -See also `T2Space.of_continuous_injective`. -/ -theorem Topology.IsEmbedding.t2Space [TopologicalSpace Y] [T2Space Y] {f : X → Y} - (hf : IsEmbedding f) : T2Space X := - .of_injective_continuous hf.injective hf.continuous - -@[deprecated (since := "2024-10-26")] -alias Embedding.t2Space := IsEmbedding.t2Space - -instance ULift.instT2Space [T2Space X] : T2Space (ULift X) := - IsEmbedding.uliftDown.t2Space - -instance [T2Space X] [TopologicalSpace Y] [T2Space Y] : - T2Space (X ⊕ Y) := by - constructor - rintro (x | x) (y | y) h - · exact separated_by_isOpenEmbedding .inl <| ne_of_apply_ne _ h - · exact separated_by_continuous continuous_isLeft <| by simp - · exact separated_by_continuous continuous_isLeft <| by simp - · exact separated_by_isOpenEmbedding .inr <| ne_of_apply_ne _ h - -instance Pi.t2Space {Y : X → Type v} [∀ a, TopologicalSpace (Y a)] - [∀ a, T2Space (Y a)] : T2Space (∀ a, Y a) := - inferInstance - -instance Sigma.t2Space {ι} {X : ι → Type*} [∀ i, TopologicalSpace (X i)] [∀ a, T2Space (X a)] : - T2Space (Σi, X i) := by - constructor - rintro ⟨i, x⟩ ⟨j, y⟩ neq - rcases eq_or_ne i j with (rfl | h) - · replace neq : x ≠ y := ne_of_apply_ne _ neq - exact separated_by_isOpenEmbedding .sigmaMk neq - · let _ := (⊥ : TopologicalSpace ι); have : DiscreteTopology ι := ⟨rfl⟩ - exact separated_by_continuous (continuous_def.2 fun u _ => isOpen_sigma_fst_preimage u) h - -section -variable (X) - -/-- The smallest equivalence relation on a topological space giving a T2 quotient. -/ -def t2Setoid : Setoid X := sInf {s | T2Space (Quotient s)} - -/-- The largest T2 quotient of a topological space. This construction is left-adjoint to the -inclusion of T2 spaces into all topological spaces. -/ -def t2Quotient := Quotient (t2Setoid X) - -namespace t2Quotient -variable {X} - -instance : TopologicalSpace (t2Quotient X) := - inferInstanceAs <| TopologicalSpace (Quotient _) - -/-- The map from a topological space to its largest T2 quotient. -/ -def mk : X → t2Quotient X := Quotient.mk (t2Setoid X) - -lemma mk_eq {x y : X} : mk x = mk y ↔ ∀ s : Setoid X, T2Space (Quotient s) → s x y := - Setoid.quotient_mk_sInf_eq - -variable (X) - -lemma surjective_mk : Surjective (mk : X → t2Quotient X) := Quotient.mk_surjective - -lemma continuous_mk : Continuous (mk : X → t2Quotient X) := - continuous_quotient_mk' - -variable {X} - -@[elab_as_elim] -protected lemma inductionOn {motive : t2Quotient X → Prop} (q : t2Quotient X) - (h : ∀ x, motive (t2Quotient.mk x)) : motive q := Quotient.inductionOn q h - -@[elab_as_elim] -protected lemma inductionOn₂ [TopologicalSpace Y] {motive : t2Quotient X → t2Quotient Y → Prop} - (q : t2Quotient X) (q' : t2Quotient Y) (h : ∀ x y, motive (mk x) (mk y)) : motive q q' := - Quotient.inductionOn₂ q q' h - -/-- The largest T2 quotient of a topological space is indeed T2. -/ -instance : T2Space (t2Quotient X) := by - rw [t2Space_iff] - rintro ⟨x⟩ ⟨y⟩ (h : ¬ t2Quotient.mk x = t2Quotient.mk y) - obtain ⟨s, hs, hsxy⟩ : ∃ s, T2Space (Quotient s) ∧ Quotient.mk s x ≠ Quotient.mk s y := by - simpa [t2Quotient.mk_eq] using h - exact separated_by_continuous (continuous_map_sInf (by exact hs)) hsxy - -lemma compatible {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] [T2Space Y] - {f : X → Y} (hf : Continuous f) : letI _ := t2Setoid X - ∀ (a b : X), a ≈ b → f a = f b := by - change t2Setoid X ≤ Setoid.ker f - exact sInf_le <| .of_injective_continuous - (Setoid.ker_lift_injective _) (hf.quotient_lift fun _ _ ↦ id) - -/-- The universal property of the largest T2 quotient of a topological space `X`: any continuous -map from `X` to a T2 space `Y` uniquely factors through `t2Quotient X`. This declaration builds the -factored map. Its continuity is `t2Quotient.continuous_lift`, the fact that it indeed factors the -original map is `t2Quotient.lift_mk` and uniquenes is `t2Quotient.unique_lift`. -/ -def lift {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] [T2Space Y] - {f : X → Y} (hf : Continuous f) : t2Quotient X → Y := - Quotient.lift f (t2Quotient.compatible hf) - -lemma continuous_lift {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] [T2Space Y] - {f : X → Y} (hf : Continuous f) : Continuous (t2Quotient.lift hf) := - continuous_coinduced_dom.mpr hf - -@[simp] -lemma lift_mk {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] [T2Space Y] - {f : X → Y} (hf : Continuous f) (x : X) : lift hf (mk x) = f x := - Quotient.lift_mk (s := t2Setoid X) f (t2Quotient.compatible hf) x - -lemma unique_lift {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] [T2Space Y] - {f : X → Y} (hf : Continuous f) {g : t2Quotient X → Y} (hfg : g ∘ mk = f) : - g = lift hf := by - apply surjective_mk X |>.right_cancellable |>.mp <| funext _ - simp [← hfg] - -end t2Quotient -end - -variable {Z : Type*} [TopologicalSpace Y] [TopologicalSpace Z] - -theorem isClosed_eq [T2Space X] {f g : Y → X} (hf : Continuous f) (hg : Continuous g) : - IsClosed { y : Y | f y = g y } := - continuous_iff_isClosed.mp (hf.prod_mk hg) _ isClosed_diagonal - -/-- If functions `f` and `g` are continuous on a closed set `s`, -then the set of points `x ∈ s` such that `f x = g x` is a closed set. -/ -protected theorem IsClosed.isClosed_eq [T2Space Y] {f g : X → Y} {s : Set X} (hs : IsClosed s) - (hf : ContinuousOn f s) (hg : ContinuousOn g s) : IsClosed {x ∈ s | f x = g x} := - (hf.prod hg).preimage_isClosed_of_isClosed hs isClosed_diagonal - -theorem isOpen_ne_fun [T2Space X] {f g : Y → X} (hf : Continuous f) (hg : Continuous g) : - IsOpen { y : Y | f y ≠ g y } := - isOpen_compl_iff.mpr <| isClosed_eq hf hg - -/-- If two continuous maps are equal on `s`, then they are equal on the closure of `s`. See also -`Set.EqOn.of_subset_closure` for a more general version. -/ -protected theorem Set.EqOn.closure [T2Space X] {s : Set Y} {f g : Y → X} (h : EqOn f g s) - (hf : Continuous f) (hg : Continuous g) : EqOn f g (closure s) := - closure_minimal h (isClosed_eq hf hg) - -/-- If two continuous functions are equal on a dense set, then they are equal. -/ -theorem Continuous.ext_on [T2Space X] {s : Set Y} (hs : Dense s) {f g : Y → X} (hf : Continuous f) - (hg : Continuous g) (h : EqOn f g s) : f = g := - funext fun x => h.closure hf hg (hs x) - -theorem eqOn_closure₂' [T2Space Z] {s : Set X} {t : Set Y} {f g : X → Y → Z} - (h : ∀ x ∈ s, ∀ y ∈ t, f x y = g x y) (hf₁ : ∀ x, Continuous (f x)) - (hf₂ : ∀ y, Continuous fun x => f x y) (hg₁ : ∀ x, Continuous (g x)) - (hg₂ : ∀ y, Continuous fun x => g x y) : ∀ x ∈ closure s, ∀ y ∈ closure t, f x y = g x y := - suffices closure s ⊆ ⋂ y ∈ closure t, { x | f x y = g x y } by simpa only [subset_def, mem_iInter] - (closure_minimal fun x hx => mem_iInter₂.2 <| Set.EqOn.closure (h x hx) (hf₁ _) (hg₁ _)) <| - isClosed_biInter fun _ _ => isClosed_eq (hf₂ _) (hg₂ _) - -theorem eqOn_closure₂ [T2Space Z] {s : Set X} {t : Set Y} {f g : X → Y → Z} - (h : ∀ x ∈ s, ∀ y ∈ t, f x y = g x y) (hf : Continuous (uncurry f)) - (hg : Continuous (uncurry g)) : ∀ x ∈ closure s, ∀ y ∈ closure t, f x y = g x y := - eqOn_closure₂' h hf.uncurry_left hf.uncurry_right hg.uncurry_left hg.uncurry_right - -/-- If `f x = g x` for all `x ∈ s` and `f`, `g` are continuous on `t`, `s ⊆ t ⊆ closure s`, then -`f x = g x` for all `x ∈ t`. See also `Set.EqOn.closure`. -/ -theorem Set.EqOn.of_subset_closure [T2Space Y] {s t : Set X} {f g : X → Y} (h : EqOn f g s) - (hf : ContinuousOn f t) (hg : ContinuousOn g t) (hst : s ⊆ t) (hts : t ⊆ closure s) : - EqOn f g t := by - intro x hx - have : (𝓝[s] x).NeBot := mem_closure_iff_clusterPt.mp (hts hx) - exact - tendsto_nhds_unique_of_eventuallyEq ((hf x hx).mono_left <| nhdsWithin_mono _ hst) - ((hg x hx).mono_left <| nhdsWithin_mono _ hst) (h.eventuallyEq_of_mem self_mem_nhdsWithin) - -theorem Function.LeftInverse.isClosed_range [T2Space X] {f : X → Y} {g : Y → X} - (h : Function.LeftInverse f g) (hf : Continuous f) (hg : Continuous g) : IsClosed (range g) := - have : EqOn (g ∘ f) id (closure <| range g) := - h.rightInvOn_range.eqOn.closure (hg.comp hf) continuous_id - isClosed_of_closure_subset fun x hx => ⟨f x, this hx⟩ - -@[deprecated (since := "2024-03-17")] -alias Function.LeftInverse.closed_range := Function.LeftInverse.isClosed_range - -theorem Function.LeftInverse.isClosedEmbedding [T2Space X] {f : X → Y} {g : Y → X} - (h : Function.LeftInverse f g) (hf : Continuous f) (hg : Continuous g) : IsClosedEmbedding g := - ⟨.of_leftInverse h hf hg, h.isClosed_range hf hg⟩ - -@[deprecated (since := "2024-10-20")] -alias Function.LeftInverse.closedEmbedding := Function.LeftInverse.isClosedEmbedding - -theorem SeparatedNhds.of_isCompact_isCompact [T2Space X] {s t : Set X} (hs : IsCompact s) - (ht : IsCompact t) (hst : Disjoint s t) : SeparatedNhds s t := by - simp only [SeparatedNhds, prod_subset_compl_diagonal_iff_disjoint.symm] at hst ⊢ - exact generalized_tube_lemma hs ht isClosed_diagonal.isOpen_compl hst - -/-- In a `T2Space X`, for disjoint closed sets `s t` such that `closure sᶜ` is compact, -there are neighbourhoods that separate `s` and `t`.-/ -lemma SeparatedNhds.of_isClosed_isCompact_closure_compl_isClosed [T2Space X] {s : Set X} - {t : Set X} (H1 : IsClosed s) (H2 : IsCompact (closure sᶜ)) (H3 : IsClosed t) - (H4 : Disjoint s t) : SeparatedNhds s t := by - -- Since `t` is a closed subset of the compact set `closure sᶜ`, it is compact. - have ht : IsCompact t := .of_isClosed_subset H2 H3 <| H4.subset_compl_left.trans subset_closure - -- we split `s` into its frontier and its interior. - rw [← diff_union_of_subset (interior_subset (s := s))] - -- since `t ⊆ sᶜ`, which is open, and `interior s` is open, we have - -- `SeparatedNhds (interior s) t`, which leaves us only with the frontier. - refine .union_left ?_ ⟨interior s, sᶜ, isOpen_interior, H1.isOpen_compl, le_rfl, - H4.subset_compl_left, disjoint_compl_right.mono_left interior_subset⟩ - -- Since the frontier of `s` is compact (as it is a subset of `closure sᶜ`), we simply apply - -- `SeparatedNhds_of_isCompact_isCompact`. - rw [← H1.frontier_eq, frontier_eq_closure_inter_closure, H1.closure_eq] - refine .of_isCompact_isCompact ?_ ht (disjoint_of_subset_left inter_subset_left H4) - exact H2.of_isClosed_subset (H1.inter isClosed_closure) inter_subset_right - -section SeparatedFinset - -theorem SeparatedNhds.of_finset_finset [T2Space X] (s t : Finset X) (h : Disjoint s t) : - SeparatedNhds (s : Set X) t := - .of_isCompact_isCompact s.finite_toSet.isCompact t.finite_toSet.isCompact <| mod_cast h - -theorem SeparatedNhds.of_singleton_finset [T2Space X] {x : X} {s : Finset X} (h : x ∉ s) : - SeparatedNhds ({x} : Set X) s := - mod_cast .of_finset_finset {x} s (Finset.disjoint_singleton_left.mpr h) - -end SeparatedFinset - -/-- In a `T2Space`, every compact set is closed. -/ -theorem IsCompact.isClosed [T2Space X] {s : Set X} (hs : IsCompact s) : IsClosed s := - isOpen_compl_iff.1 <| isOpen_iff_forall_mem_open.mpr fun x hx => - let ⟨u, v, _, vo, su, xv, uv⟩ := - SeparatedNhds.of_isCompact_isCompact hs isCompact_singleton (disjoint_singleton_right.2 hx) - ⟨v, (uv.mono_left <| show s ≤ u from su).subset_compl_left, vo, by simpa using xv⟩ - -theorem IsCompact.preimage_continuous [CompactSpace X] [T2Space Y] {f : X → Y} {s : Set Y} - (hs : IsCompact s) (hf : Continuous f) : IsCompact (f ⁻¹' s) := - (hs.isClosed.preimage hf).isCompact - -lemma Pi.isCompact_iff {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] - [∀ i, T2Space (π i)] {s : Set (Π i, π i)} : - IsCompact s ↔ IsClosed s ∧ ∀ i, IsCompact (eval i '' s) := by - constructor <;> intro H - · exact ⟨H.isClosed, fun i ↦ H.image <| continuous_apply i⟩ - · exact IsCompact.of_isClosed_subset (isCompact_univ_pi H.2) H.1 (subset_pi_eval_image univ s) - -lemma Pi.isCompact_closure_iff {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] - [∀ i, T2Space (π i)] {s : Set (Π i, π i)} : - IsCompact (closure s) ↔ ∀ i, IsCompact (closure <| eval i '' s) := by - simp_rw [← exists_isCompact_superset_iff, Pi.exists_compact_superset_iff, image_subset_iff] - -/-- If `V : ι → Set X` is a decreasing family of compact sets then any neighborhood of -`⋂ i, V i` contains some `V i`. This is a version of `exists_subset_nhds_of_isCompact'` where we -don't need to assume each `V i` closed because it follows from compactness since `X` is -assumed to be Hausdorff. -/ -theorem exists_subset_nhds_of_isCompact [T2Space X] {ι : Type*} [Nonempty ι] {V : ι → Set X} - (hV : Directed (· ⊇ ·) V) (hV_cpct : ∀ i, IsCompact (V i)) {U : Set X} - (hU : ∀ x ∈ ⋂ i, V i, U ∈ 𝓝 x) : ∃ i, V i ⊆ U := - exists_subset_nhds_of_isCompact' hV hV_cpct (fun i => (hV_cpct i).isClosed) hU - -theorem CompactExhaustion.isClosed [T2Space X] (K : CompactExhaustion X) (n : ℕ) : IsClosed (K n) := - (K.isCompact n).isClosed - -theorem IsCompact.inter [T2Space X] {s t : Set X} (hs : IsCompact s) (ht : IsCompact t) : - IsCompact (s ∩ t) := - hs.inter_right <| ht.isClosed - -theorem image_closure_of_isCompact [T2Space Y] {s : Set X} (hs : IsCompact (closure s)) {f : X → Y} - (hf : ContinuousOn f (closure s)) : f '' closure s = closure (f '' s) := - Subset.antisymm hf.image_closure <| - closure_minimal (image_subset f subset_closure) (hs.image_of_continuousOn hf).isClosed - -/-- A continuous map from a compact space to a Hausdorff space is a closed map. -/ -protected theorem Continuous.isClosedMap [CompactSpace X] [T2Space Y] {f : X → Y} - (h : Continuous f) : IsClosedMap f := fun _s hs => (hs.isCompact.image h).isClosed - -/-- A continuous injective map from a compact space to a Hausdorff space is a closed embedding. -/ -theorem Continuous.isClosedEmbedding [CompactSpace X] [T2Space Y] {f : X → Y} (h : Continuous f) - (hf : Function.Injective f) : IsClosedEmbedding f := - .of_continuous_injective_isClosedMap h hf h.isClosedMap - -@[deprecated (since := "2024-10-20")] -alias Continuous.closedEmbedding := Continuous.isClosedEmbedding - -/-- A continuous surjective map from a compact space to a Hausdorff space is a quotient map. -/ -theorem IsQuotientMap.of_surjective_continuous [CompactSpace X] [T2Space Y] {f : X → Y} - (hsurj : Surjective f) (hcont : Continuous f) : IsQuotientMap f := - hcont.isClosedMap.isQuotientMap hcont hsurj - -@[deprecated (since := "2024-10-22")] -alias QuotientMap.of_surjective_continuous := IsQuotientMap.of_surjective_continuous - -theorem isPreirreducible_iff_subsingleton [T2Space X] {S : Set X} : - IsPreirreducible S ↔ S.Subsingleton := by - refine ⟨fun h x hx y hy => ?_, Set.Subsingleton.isPreirreducible⟩ - by_contra e - obtain ⟨U, V, hU, hV, hxU, hyV, h'⟩ := t2_separation e - exact ((h U V hU hV ⟨x, hx, hxU⟩ ⟨y, hy, hyV⟩).mono inter_subset_right).not_disjoint h' - --- todo: use `alias` + `attribute [protected]` once we get `attribute [protected]` -protected lemma IsPreirreducible.subsingleton [T2Space X] {S : Set X} (h : IsPreirreducible S) : - S.Subsingleton := - isPreirreducible_iff_subsingleton.1 h - -theorem isIrreducible_iff_singleton [T2Space X] {S : Set X} : IsIrreducible S ↔ ∃ x, S = {x} := by - rw [IsIrreducible, isPreirreducible_iff_subsingleton, - exists_eq_singleton_iff_nonempty_subsingleton] - -/-- There does not exist a nontrivial preirreducible T₂ space. -/ -theorem not_preirreducible_nontrivial_t2 (X) [TopologicalSpace X] [PreirreducibleSpace X] - [Nontrivial X] [T2Space X] : False := - (PreirreducibleSpace.isPreirreducible_univ (X := X)).subsingleton.not_nontrivial nontrivial_univ - -theorem t2Space_antitone {X : Type*} : Antitone (@T2Space X) := - fun inst₁ inst₂ h_top h_t2 ↦ @T2Space.of_injective_continuous _ _ inst₁ inst₂ - h_t2 _ Function.injective_id <| continuous_id_of_le h_top - end Separation - -section RegularSpace - -/-- A topological space is called a *regular space* if for any closed set `s` and `a ∉ s`, there -exist disjoint open sets `U ⊇ s` and `V ∋ a`. We formulate this condition in terms of `Disjoint`ness -of filters `𝓝ˢ s` and `𝓝 a`. -/ -@[mk_iff] -class RegularSpace (X : Type u) [TopologicalSpace X] : Prop where - /-- If `a` is a point that does not belong to a closed set `s`, then `a` and `s` admit disjoint - neighborhoods. -/ - regular : ∀ {s : Set X} {a}, IsClosed s → a ∉ s → Disjoint (𝓝ˢ s) (𝓝 a) - -theorem regularSpace_TFAE (X : Type u) [TopologicalSpace X] : - List.TFAE [RegularSpace X, - ∀ (s : Set X) x, x ∉ closure s → Disjoint (𝓝ˢ s) (𝓝 x), - ∀ (x : X) (s : Set X), Disjoint (𝓝ˢ s) (𝓝 x) ↔ x ∉ closure s, - ∀ (x : X) (s : Set X), s ∈ 𝓝 x → ∃ t ∈ 𝓝 x, IsClosed t ∧ t ⊆ s, - ∀ x : X, (𝓝 x).lift' closure ≤ 𝓝 x, - ∀ x : X , (𝓝 x).lift' closure = 𝓝 x] := by - tfae_have 1 ↔ 5 := by - rw [regularSpace_iff, (@compl_surjective (Set X) _).forall, forall_swap] - simp only [isClosed_compl_iff, mem_compl_iff, Classical.not_not, @and_comm (_ ∈ _), - (nhds_basis_opens _).lift'_closure.le_basis_iff (nhds_basis_opens _), and_imp, - (nhds_basis_opens _).disjoint_iff_right, exists_prop, ← subset_interior_iff_mem_nhdsSet, - interior_compl, compl_subset_compl] - tfae_have 5 → 6 := fun h a => (h a).antisymm (𝓝 _).le_lift'_closure - tfae_have 6 → 4 - | H, a, s, hs => by - rw [← H] at hs - rcases (𝓝 a).basis_sets.lift'_closure.mem_iff.mp hs with ⟨U, hU, hUs⟩ - exact ⟨closure U, mem_of_superset hU subset_closure, isClosed_closure, hUs⟩ - tfae_have 4 → 2 - | H, s, a, ha => by - have ha' : sᶜ ∈ 𝓝 a := by rwa [← mem_interior_iff_mem_nhds, interior_compl] - rcases H _ _ ha' with ⟨U, hU, hUc, hUs⟩ - refine disjoint_of_disjoint_of_mem disjoint_compl_left ?_ hU - rwa [← subset_interior_iff_mem_nhdsSet, hUc.isOpen_compl.interior_eq, subset_compl_comm] - tfae_have 2 → 3 := by - refine fun H a s => ⟨fun hd has => mem_closure_iff_nhds_ne_bot.mp has ?_, H s a⟩ - exact (hd.symm.mono_right <| @principal_le_nhdsSet _ _ s).eq_bot - tfae_have 3 → 1 := fun H => ⟨fun hs ha => (H _ _).mpr <| hs.closure_eq.symm ▸ ha⟩ - tfae_finish - -theorem RegularSpace.of_lift'_closure_le (h : ∀ x : X, (𝓝 x).lift' closure ≤ 𝓝 x) : - RegularSpace X := - Iff.mpr ((regularSpace_TFAE X).out 0 4) h - -theorem RegularSpace.of_lift'_closure (h : ∀ x : X, (𝓝 x).lift' closure = 𝓝 x) : RegularSpace X := - Iff.mpr ((regularSpace_TFAE X).out 0 5) h - -@[deprecated (since := "2024-02-28")] -alias RegularSpace.ofLift'_closure := RegularSpace.of_lift'_closure - -theorem RegularSpace.of_hasBasis {ι : X → Sort*} {p : ∀ a, ι a → Prop} {s : ∀ a, ι a → Set X} - (h₁ : ∀ a, (𝓝 a).HasBasis (p a) (s a)) (h₂ : ∀ a i, p a i → IsClosed (s a i)) : - RegularSpace X := - .of_lift'_closure fun a => (h₁ a).lift'_closure_eq_self (h₂ a) - -@[deprecated (since := "2024-02-28")] -alias RegularSpace.ofBasis := RegularSpace.of_hasBasis - -theorem RegularSpace.of_exists_mem_nhds_isClosed_subset - (h : ∀ (x : X), ∀ s ∈ 𝓝 x, ∃ t ∈ 𝓝 x, IsClosed t ∧ t ⊆ s) : RegularSpace X := - Iff.mpr ((regularSpace_TFAE X).out 0 3) h - -@[deprecated (since := "2024-02-28")] -alias RegularSpace.ofExistsMemNhdsIsClosedSubset := RegularSpace.of_exists_mem_nhds_isClosed_subset - -/-- A weakly locally compact R₁ space is regular. -/ -instance (priority := 100) [WeaklyLocallyCompactSpace X] [R1Space X] : RegularSpace X := - .of_hasBasis isCompact_isClosed_basis_nhds fun _ _ ⟨_, _, h⟩ ↦ h - -section -variable [RegularSpace X] {x : X} {s : Set X} - -theorem disjoint_nhdsSet_nhds : Disjoint (𝓝ˢ s) (𝓝 x) ↔ x ∉ closure s := by - have h := (regularSpace_TFAE X).out 0 2 - exact h.mp ‹_› _ _ - -theorem disjoint_nhds_nhdsSet : Disjoint (𝓝 x) (𝓝ˢ s) ↔ x ∉ closure s := - disjoint_comm.trans disjoint_nhdsSet_nhds - -/-- A regular space is R₁. -/ -instance (priority := 100) : R1Space X where - specializes_or_disjoint_nhds _ _ := or_iff_not_imp_left.2 fun h ↦ by - rwa [← nhdsSet_singleton, disjoint_nhdsSet_nhds, ← specializes_iff_mem_closure] - -theorem exists_mem_nhds_isClosed_subset {x : X} {s : Set X} (h : s ∈ 𝓝 x) : - ∃ t ∈ 𝓝 x, IsClosed t ∧ t ⊆ s := by - have h' := (regularSpace_TFAE X).out 0 3 - exact h'.mp ‹_› _ _ h - -theorem closed_nhds_basis (x : X) : (𝓝 x).HasBasis (fun s : Set X => s ∈ 𝓝 x ∧ IsClosed s) id := - hasBasis_self.2 fun _ => exists_mem_nhds_isClosed_subset - -theorem lift'_nhds_closure (x : X) : (𝓝 x).lift' closure = 𝓝 x := - (closed_nhds_basis x).lift'_closure_eq_self fun _ => And.right - -theorem Filter.HasBasis.nhds_closure {ι : Sort*} {x : X} {p : ι → Prop} {s : ι → Set X} - (h : (𝓝 x).HasBasis p s) : (𝓝 x).HasBasis p fun i => closure (s i) := - lift'_nhds_closure x ▸ h.lift'_closure - -theorem hasBasis_nhds_closure (x : X) : (𝓝 x).HasBasis (fun s => s ∈ 𝓝 x) closure := - (𝓝 x).basis_sets.nhds_closure - -theorem hasBasis_opens_closure (x : X) : (𝓝 x).HasBasis (fun s => x ∈ s ∧ IsOpen s) closure := - (nhds_basis_opens x).nhds_closure - -theorem IsCompact.exists_isOpen_closure_subset {K U : Set X} (hK : IsCompact K) (hU : U ∈ 𝓝ˢ K) : - ∃ V, IsOpen V ∧ K ⊆ V ∧ closure V ⊆ U := by - have hd : Disjoint (𝓝ˢ K) (𝓝ˢ Uᶜ) := by - simpa [hK.disjoint_nhdsSet_left, disjoint_nhds_nhdsSet, - ← subset_interior_iff_mem_nhdsSet] using hU - rcases ((hasBasis_nhdsSet _).disjoint_iff (hasBasis_nhdsSet _)).1 hd - with ⟨V, ⟨hVo, hKV⟩, W, ⟨hW, hUW⟩, hVW⟩ - refine ⟨V, hVo, hKV, Subset.trans ?_ (compl_subset_comm.1 hUW)⟩ - exact closure_minimal hVW.subset_compl_right hW.isClosed_compl - -theorem IsCompact.lift'_closure_nhdsSet {K : Set X} (hK : IsCompact K) : - (𝓝ˢ K).lift' closure = 𝓝ˢ K := by - refine le_antisymm (fun U hU ↦ ?_) (le_lift'_closure _) - rcases hK.exists_isOpen_closure_subset hU with ⟨V, hVo, hKV, hVU⟩ - exact mem_of_superset (mem_lift' <| hVo.mem_nhdsSet.2 hKV) hVU - -theorem TopologicalSpace.IsTopologicalBasis.nhds_basis_closure {B : Set (Set X)} - (hB : IsTopologicalBasis B) (x : X) : - (𝓝 x).HasBasis (fun s : Set X => x ∈ s ∧ s ∈ B) closure := by - simpa only [and_comm] using hB.nhds_hasBasis.nhds_closure - -theorem TopologicalSpace.IsTopologicalBasis.exists_closure_subset {B : Set (Set X)} - (hB : IsTopologicalBasis B) {x : X} {s : Set X} (h : s ∈ 𝓝 x) : - ∃ t ∈ B, x ∈ t ∧ closure t ⊆ s := by - simpa only [exists_prop, and_assoc] using hB.nhds_hasBasis.nhds_closure.mem_iff.mp h - -protected theorem Topology.IsInducing.regularSpace [TopologicalSpace Y] {f : Y → X} - (hf : IsInducing f) : RegularSpace Y := - .of_hasBasis - (fun b => by rw [hf.nhds_eq_comap b]; exact (closed_nhds_basis _).comap _) - fun b s hs => by exact hs.2.preimage hf.continuous - -@[deprecated (since := "2024-10-28")] alias Inducing.regularSpace := IsInducing.regularSpace - -theorem regularSpace_induced (f : Y → X) : @RegularSpace Y (induced f ‹_›) := - letI := induced f ‹_› - (IsInducing.induced f).regularSpace - -theorem regularSpace_sInf {X} {T : Set (TopologicalSpace X)} (h : ∀ t ∈ T, @RegularSpace X t) : - @RegularSpace X (sInf T) := by - let _ := sInf T - have : ∀ a, (𝓝 a).HasBasis - (fun If : Σ I : Set T, I → Set X => - If.1.Finite ∧ ∀ i : If.1, If.2 i ∈ @nhds X i a ∧ @IsClosed X i (If.2 i)) - fun If => ⋂ i : If.1, If.snd i := fun a ↦ by - rw [nhds_sInf, ← iInf_subtype''] - exact hasBasis_iInf fun t : T => @closed_nhds_basis X t (h t t.2) a - refine .of_hasBasis this fun a If hIf => isClosed_iInter fun i => ?_ - exact (hIf.2 i).2.mono (sInf_le (i : T).2) - -theorem regularSpace_iInf {ι X} {t : ι → TopologicalSpace X} (h : ∀ i, @RegularSpace X (t i)) : - @RegularSpace X (iInf t) := - regularSpace_sInf <| forall_mem_range.mpr h - -theorem RegularSpace.inf {X} {t₁ t₂ : TopologicalSpace X} (h₁ : @RegularSpace X t₁) - (h₂ : @RegularSpace X t₂) : @RegularSpace X (t₁ ⊓ t₂) := by - rw [inf_eq_iInf] - exact regularSpace_iInf (Bool.forall_bool.2 ⟨h₂, h₁⟩) - -instance {p : X → Prop} : RegularSpace (Subtype p) := - IsEmbedding.subtypeVal.isInducing.regularSpace - -instance [TopologicalSpace Y] [RegularSpace Y] : RegularSpace (X × Y) := - (regularSpace_induced (@Prod.fst X Y)).inf (regularSpace_induced (@Prod.snd X Y)) - -instance {ι : Type*} {X : ι → Type*} [∀ i, TopologicalSpace (X i)] [∀ i, RegularSpace (X i)] : - RegularSpace (∀ i, X i) := - regularSpace_iInf fun _ => regularSpace_induced _ - -/-- In a regular space, if a compact set and a closed set are disjoint, then they have disjoint -neighborhoods. -/ -lemma SeparatedNhds.of_isCompact_isClosed {s t : Set X} - (hs : IsCompact s) (ht : IsClosed t) (hst : Disjoint s t) : SeparatedNhds s t := by - simpa only [separatedNhds_iff_disjoint, hs.disjoint_nhdsSet_left, disjoint_nhds_nhdsSet, - ht.closure_eq, disjoint_left] using hst - -end - -/-- This technique to witness `HasSeparatingCover` in regular Lindelöf topological spaces -will be used to prove regular Lindelöf spaces are normal. -/ -lemma IsClosed.HasSeparatingCover {s t : Set X} [LindelofSpace X] [RegularSpace X] - (s_cl : IsClosed s) (t_cl : IsClosed t) (st_dis : Disjoint s t) : HasSeparatingCover s t := by - -- `IsLindelof.indexed_countable_subcover` requires the space be Nonempty - rcases isEmpty_or_nonempty X with empty_X | nonempty_X - · rw [subset_eq_empty (t := s) (fun ⦃_⦄ _ ↦ trivial) (univ_eq_empty_iff.mpr empty_X)] - exact hasSeparatingCovers_iff_separatedNhds.mpr (SeparatedNhds.empty_left t) |>.1 - -- This is almost `HasSeparatingCover`, but is not countable. We define for all `a : X` for use - -- with `IsLindelof.indexed_countable_subcover` momentarily. - have (a : X) : ∃ n : Set X, IsOpen n ∧ Disjoint (closure n) t ∧ (a ∈ s → a ∈ n) := by - wlog ains : a ∈ s - · exact ⟨∅, isOpen_empty, SeparatedNhds.empty_left t |>.disjoint_closure_left, fun a ↦ ains a⟩ - obtain ⟨n, nna, ncl, nsubkc⟩ := ((regularSpace_TFAE X).out 0 3 :).mp ‹RegularSpace X› a tᶜ <| - t_cl.compl_mem_nhds (disjoint_left.mp st_dis ains) - exact - ⟨interior n, - isOpen_interior, - disjoint_left.mpr fun ⦃_⦄ ain ↦ - nsubkc <| (IsClosed.closure_subset_iff ncl).mpr interior_subset ain, - fun _ ↦ mem_interior_iff_mem_nhds.mpr nna⟩ - -- By Lindelöf, we may obtain a countable subcover witnessing `HasSeparatingCover` - choose u u_open u_dis u_nhd using this - obtain ⟨f, f_cov⟩ := s_cl.isLindelof.indexed_countable_subcover - u u_open (fun a ainh ↦ mem_iUnion.mpr ⟨a, u_nhd a ainh⟩) - exact ⟨u ∘ f, f_cov, fun n ↦ ⟨u_open (f n), u_dis (f n)⟩⟩ - - -end RegularSpace - -section LocallyCompactRegularSpace - -/-- In a (possibly non-Hausdorff) locally compact regular space, for every containment `K ⊆ U` of - a compact set `K` in an open set `U`, there is a compact closed neighborhood `L` - such that `K ⊆ L ⊆ U`: equivalently, there is a compact closed set `L` such - that `K ⊆ interior L` and `L ⊆ U`. -/ -theorem exists_compact_closed_between [LocallyCompactSpace X] [RegularSpace X] - {K U : Set X} (hK : IsCompact K) (hU : IsOpen U) (h_KU : K ⊆ U) : - ∃ L, IsCompact L ∧ IsClosed L ∧ K ⊆ interior L ∧ L ⊆ U := - let ⟨L, L_comp, KL, LU⟩ := exists_compact_between hK hU h_KU - ⟨closure L, L_comp.closure, isClosed_closure, KL.trans <| interior_mono subset_closure, - L_comp.closure_subset_of_isOpen hU LU⟩ - -/-- In a locally compact regular space, given a compact set `K` inside an open set `U`, we can find -an open set `V` between these sets with compact closure: `K ⊆ V` and the closure of `V` is -inside `U`. -/ -theorem exists_open_between_and_isCompact_closure [LocallyCompactSpace X] [RegularSpace X] - {K U : Set X} (hK : IsCompact K) (hU : IsOpen U) (hKU : K ⊆ U) : - ∃ V, IsOpen V ∧ K ⊆ V ∧ closure V ⊆ U ∧ IsCompact (closure V) := by - rcases exists_compact_closed_between hK hU hKU with ⟨L, L_compact, L_closed, KL, LU⟩ - have A : closure (interior L) ⊆ L := by - apply (closure_mono interior_subset).trans (le_of_eq L_closed.closure_eq) - refine ⟨interior L, isOpen_interior, KL, A.trans LU, ?_⟩ - exact L_compact.closure_of_subset interior_subset - -end LocallyCompactRegularSpace - -section T25 - -/-- A T₂.₅ space, also known as a Urysohn space, is a topological space - where for every pair `x ≠ y`, there are two open sets, with the intersection of closures - empty, one containing `x` and the other `y` . -/ -class T25Space (X : Type u) [TopologicalSpace X] : Prop where - /-- Given two distinct points in a T₂.₅ space, their filters of closed neighborhoods are - disjoint. -/ - t2_5 : ∀ ⦃x y : X⦄, x ≠ y → Disjoint ((𝓝 x).lift' closure) ((𝓝 y).lift' closure) - -@[simp] -theorem disjoint_lift'_closure_nhds [T25Space X] {x y : X} : - Disjoint ((𝓝 x).lift' closure) ((𝓝 y).lift' closure) ↔ x ≠ y := - ⟨fun h hxy => by simp [hxy, nhds_neBot.ne] at h, fun h => T25Space.t2_5 h⟩ - --- see Note [lower instance priority] -instance (priority := 100) T25Space.t2Space [T25Space X] : T2Space X := - t2Space_iff_disjoint_nhds.2 fun _ _ hne => - (disjoint_lift'_closure_nhds.2 hne).mono (le_lift'_closure _) (le_lift'_closure _) - -theorem exists_nhds_disjoint_closure [T25Space X] {x y : X} (h : x ≠ y) : - ∃ s ∈ 𝓝 x, ∃ t ∈ 𝓝 y, Disjoint (closure s) (closure t) := - ((𝓝 x).basis_sets.lift'_closure.disjoint_iff (𝓝 y).basis_sets.lift'_closure).1 <| - disjoint_lift'_closure_nhds.2 h - -theorem exists_open_nhds_disjoint_closure [T25Space X] {x y : X} (h : x ≠ y) : - ∃ u : Set X, - x ∈ u ∧ IsOpen u ∧ ∃ v : Set X, y ∈ v ∧ IsOpen v ∧ Disjoint (closure u) (closure v) := by - simpa only [exists_prop, and_assoc] using - ((nhds_basis_opens x).lift'_closure.disjoint_iff (nhds_basis_opens y).lift'_closure).1 - (disjoint_lift'_closure_nhds.2 h) - -theorem T25Space.of_injective_continuous [TopologicalSpace Y] [T25Space Y] {f : X → Y} - (hinj : Injective f) (hcont : Continuous f) : T25Space X where - t2_5 x y hne := (tendsto_lift'_closure_nhds hcont x).disjoint (t2_5 <| hinj.ne hne) - (tendsto_lift'_closure_nhds hcont y) - -theorem Topology.IsEmbedding.t25Space [TopologicalSpace Y] [T25Space Y] {f : X → Y} - (hf : IsEmbedding f) : T25Space X := - .of_injective_continuous hf.injective hf.continuous - -@[deprecated (since := "2024-10-26")] -alias Embedding.t25Space := IsEmbedding.t25Space - -instance Subtype.instT25Space [T25Space X] {p : X → Prop} : T25Space {x // p x} := - IsEmbedding.subtypeVal.t25Space - -end T25 - -section T3 - -/-- A T₃ space is a T₀ space which is a regular space. Any T₃ space is a T₁ space, a T₂ space, and -a T₂.₅ space. -/ -class T3Space (X : Type u) [TopologicalSpace X] extends T0Space X, RegularSpace X : Prop - -instance (priority := 90) instT3Space [T0Space X] [RegularSpace X] : T3Space X := ⟨⟩ - -theorem RegularSpace.t3Space_iff_t0Space [RegularSpace X] : T3Space X ↔ T0Space X := by - constructor <;> intro <;> infer_instance - --- see Note [lower instance priority] -instance (priority := 100) T3Space.t25Space [T3Space X] : T25Space X := by - refine ⟨fun x y hne => ?_⟩ - rw [lift'_nhds_closure, lift'_nhds_closure] - have : x ∉ closure {y} ∨ y ∉ closure {x} := - (t0Space_iff_or_not_mem_closure X).mp inferInstance hne - simp only [← disjoint_nhds_nhdsSet, nhdsSet_singleton] at this - exact this.elim id fun h => h.symm - -protected theorem Topology.IsEmbedding.t3Space [TopologicalSpace Y] [T3Space Y] {f : X → Y} - (hf : IsEmbedding f) : T3Space X := - { toT0Space := hf.t0Space - toRegularSpace := hf.isInducing.regularSpace } - -@[deprecated (since := "2024-10-26")] -alias Embedding.t3Space := IsEmbedding.t3Space - -instance Subtype.t3Space [T3Space X] {p : X → Prop} : T3Space (Subtype p) := - IsEmbedding.subtypeVal.t3Space - -instance ULift.instT3Space [T3Space X] : T3Space (ULift X) := - IsEmbedding.uliftDown.t3Space - -instance [TopologicalSpace Y] [T3Space X] [T3Space Y] : T3Space (X × Y) := ⟨⟩ - -instance {ι : Type*} {X : ι → Type*} [∀ i, TopologicalSpace (X i)] [∀ i, T3Space (X i)] : - T3Space (∀ i, X i) := ⟨⟩ - -/-- Given two points `x ≠ y`, we can find neighbourhoods `x ∈ V₁ ⊆ U₁` and `y ∈ V₂ ⊆ U₂`, -with the `Vₖ` closed and the `Uₖ` open, such that the `Uₖ` are disjoint. -/ -theorem disjoint_nested_nhds [T3Space X] {x y : X} (h : x ≠ y) : - ∃ U₁ ∈ 𝓝 x, ∃ V₁ ∈ 𝓝 x, ∃ U₂ ∈ 𝓝 y, ∃ V₂ ∈ 𝓝 y, - IsClosed V₁ ∧ IsClosed V₂ ∧ IsOpen U₁ ∧ IsOpen U₂ ∧ V₁ ⊆ U₁ ∧ V₂ ⊆ U₂ ∧ Disjoint U₁ U₂ := by - rcases t2_separation h with ⟨U₁, U₂, U₁_op, U₂_op, x_in, y_in, H⟩ - rcases exists_mem_nhds_isClosed_subset (U₁_op.mem_nhds x_in) with ⟨V₁, V₁_in, V₁_closed, h₁⟩ - rcases exists_mem_nhds_isClosed_subset (U₂_op.mem_nhds y_in) with ⟨V₂, V₂_in, V₂_closed, h₂⟩ - exact ⟨U₁, mem_of_superset V₁_in h₁, V₁, V₁_in, U₂, mem_of_superset V₂_in h₂, V₂, V₂_in, - V₁_closed, V₂_closed, U₁_op, U₂_op, h₁, h₂, H⟩ - -open SeparationQuotient - -/-- The `SeparationQuotient` of a regular space is a T₃ space. -/ -instance [RegularSpace X] : T3Space (SeparationQuotient X) where - regular {s a} hs ha := by - rcases surjective_mk a with ⟨a, rfl⟩ - rw [← disjoint_comap_iff surjective_mk, comap_mk_nhds_mk, comap_mk_nhdsSet] - exact RegularSpace.regular (hs.preimage continuous_mk) ha - -end T3 - -section NormalSpace - -/-- A topological space is said to be a *normal space* if any two disjoint closed sets -have disjoint open neighborhoods. -/ -class NormalSpace (X : Type u) [TopologicalSpace X] : Prop where - /-- Two disjoint sets in a normal space admit disjoint neighbourhoods. -/ - normal : ∀ s t : Set X, IsClosed s → IsClosed t → Disjoint s t → SeparatedNhds s t - -theorem normal_separation [NormalSpace X] {s t : Set X} (H1 : IsClosed s) (H2 : IsClosed t) - (H3 : Disjoint s t) : SeparatedNhds s t := - NormalSpace.normal s t H1 H2 H3 - -theorem disjoint_nhdsSet_nhdsSet [NormalSpace X] {s t : Set X} (hs : IsClosed s) (ht : IsClosed t) - (hd : Disjoint s t) : Disjoint (𝓝ˢ s) (𝓝ˢ t) := - (normal_separation hs ht hd).disjoint_nhdsSet - -theorem normal_exists_closure_subset [NormalSpace X] {s t : Set X} (hs : IsClosed s) (ht : IsOpen t) - (hst : s ⊆ t) : ∃ u, IsOpen u ∧ s ⊆ u ∧ closure u ⊆ t := by - have : Disjoint s tᶜ := Set.disjoint_left.mpr fun x hxs hxt => hxt (hst hxs) - rcases normal_separation hs (isClosed_compl_iff.2 ht) this with - ⟨s', t', hs', ht', hss', htt', hs't'⟩ - refine ⟨s', hs', hss', Subset.trans (closure_minimal ?_ (isClosed_compl_iff.2 ht')) - (compl_subset_comm.1 htt')⟩ - exact fun x hxs hxt => hs't'.le_bot ⟨hxs, hxt⟩ - -/-- If the codomain of a closed embedding is a normal space, then so is the domain. -/ -protected theorem Topology.IsClosedEmbedding.normalSpace [TopologicalSpace Y] [NormalSpace Y] - {f : X → Y} (hf : IsClosedEmbedding f) : NormalSpace X where - normal s t hs ht hst := by - have H : SeparatedNhds (f '' s) (f '' t) := - NormalSpace.normal (f '' s) (f '' t) (hf.isClosedMap s hs) (hf.isClosedMap t ht) - (disjoint_image_of_injective hf.injective hst) - exact (H.preimage hf.continuous).mono (subset_preimage_image _ _) (subset_preimage_image _ _) - -@[deprecated (since := "2024-10-20")] -alias ClosedEmbedding.normalSpace := IsClosedEmbedding.normalSpace - -instance (priority := 100) NormalSpace.of_compactSpace_r1Space [CompactSpace X] [R1Space X] : - NormalSpace X where - normal _s _t hs ht := .of_isCompact_isCompact_isClosed hs.isCompact ht.isCompact ht - -set_option pp.universes true in -/-- A regular topological space with a Lindelöf topology is a normal space. A consequence of e.g. -Corollaries 20.8 and 20.10 of [Willard's *General Topology*][zbMATH02107988] (without the -assumption of Hausdorff). -/ -instance (priority := 100) NormalSpace.of_regularSpace_lindelofSpace - [RegularSpace X] [LindelofSpace X] : NormalSpace X where - normal _ _ hcl kcl hkdis := - hasSeparatingCovers_iff_separatedNhds.mp - ⟨hcl.HasSeparatingCover kcl hkdis, kcl.HasSeparatingCover hcl (Disjoint.symm hkdis)⟩ - -instance (priority := 100) NormalSpace.of_regularSpace_secondCountableTopology - [RegularSpace X] [SecondCountableTopology X] : NormalSpace X := - of_regularSpace_lindelofSpace - -end NormalSpace - -section Normality - -/-- A T₄ space is a normal T₁ space. -/ -class T4Space (X : Type u) [TopologicalSpace X] extends T1Space X, NormalSpace X : Prop - -instance (priority := 100) [T1Space X] [NormalSpace X] : T4Space X := ⟨⟩ - --- see Note [lower instance priority] -instance (priority := 100) T4Space.t3Space [T4Space X] : T3Space X where - regular hs hxs := by simpa only [nhdsSet_singleton] using (normal_separation hs isClosed_singleton - (disjoint_singleton_right.mpr hxs)).disjoint_nhdsSet - -/-- If the codomain of a closed embedding is a T₄ space, then so is the domain. -/ -protected theorem Topology.IsClosedEmbedding.t4Space [TopologicalSpace Y] [T4Space Y] {f : X → Y} - (hf : IsClosedEmbedding f) : T4Space X where - toT1Space := hf.isEmbedding.t1Space - toNormalSpace := hf.normalSpace - -@[deprecated (since := "2024-10-20")] -alias ClosedEmbedding.t4Space := IsClosedEmbedding.t4Space - -instance ULift.instT4Space [T4Space X] : T4Space (ULift X) := IsClosedEmbedding.uliftDown.t4Space - -namespace SeparationQuotient - -/-- The `SeparationQuotient` of a normal space is a normal space. -/ -instance [NormalSpace X] : NormalSpace (SeparationQuotient X) where - normal s t hs ht hd := separatedNhds_iff_disjoint.2 <| by - rw [← disjoint_comap_iff surjective_mk, comap_mk_nhdsSet, comap_mk_nhdsSet] - exact disjoint_nhdsSet_nhdsSet (hs.preimage continuous_mk) (ht.preimage continuous_mk) - (hd.preimage mk) - -end SeparationQuotient - -variable (X) - -end Normality - -section CompletelyNormal - -/-- A topological space `X` is a *completely normal space* provided that for any two sets `s`, `t` -such that if both `closure s` is disjoint with `t`, and `s` is disjoint with `closure t`, -then there exist disjoint neighbourhoods of `s` and `t`. -/ -class CompletelyNormalSpace (X : Type u) [TopologicalSpace X] : Prop where - /-- If `closure s` is disjoint with `t`, and `s` is disjoint with `closure t`, then `s` and `t` - admit disjoint neighbourhoods. -/ - completely_normal : - ∀ ⦃s t : Set X⦄, Disjoint (closure s) t → Disjoint s (closure t) → Disjoint (𝓝ˢ s) (𝓝ˢ t) - -export CompletelyNormalSpace (completely_normal) - --- see Note [lower instance priority] -/-- A completely normal space is a normal space. -/ -instance (priority := 100) CompletelyNormalSpace.toNormalSpace - [CompletelyNormalSpace X] : NormalSpace X where - normal s t hs ht hd := separatedNhds_iff_disjoint.2 <| - completely_normal (by rwa [hs.closure_eq]) (by rwa [ht.closure_eq]) - -theorem Topology.IsEmbedding.completelyNormalSpace [TopologicalSpace Y] [CompletelyNormalSpace Y] - {e : X → Y} (he : IsEmbedding e) : CompletelyNormalSpace X := by - refine ⟨fun s t hd₁ hd₂ => ?_⟩ - simp only [he.isInducing.nhdsSet_eq_comap] - refine disjoint_comap (completely_normal ?_ ?_) - · rwa [← subset_compl_iff_disjoint_left, image_subset_iff, preimage_compl, - ← he.closure_eq_preimage_closure_image, subset_compl_iff_disjoint_left] - · rwa [← subset_compl_iff_disjoint_right, image_subset_iff, preimage_compl, - ← he.closure_eq_preimage_closure_image, subset_compl_iff_disjoint_right] - -@[deprecated (since := "2024-10-26")] -alias Embedding.completelyNormalSpace := IsEmbedding.completelyNormalSpace - -/-- A subspace of a completely normal space is a completely normal space. -/ -instance [CompletelyNormalSpace X] {p : X → Prop} : CompletelyNormalSpace { x // p x } := - IsEmbedding.subtypeVal.completelyNormalSpace - -instance ULift.instCompletelyNormalSpace [CompletelyNormalSpace X] : - CompletelyNormalSpace (ULift X) := - IsEmbedding.uliftDown.completelyNormalSpace - -/-- A T₅ space is a completely normal T₁ space. -/ -class T5Space (X : Type u) [TopologicalSpace X] extends T1Space X, CompletelyNormalSpace X : Prop - -theorem Topology.IsEmbedding.t5Space [TopologicalSpace Y] [T5Space Y] {e : X → Y} - (he : IsEmbedding e) : T5Space X where - __ := he.t1Space - completely_normal := by - have := he.completelyNormalSpace - exact completely_normal - -@[deprecated (since := "2024-10-26")] -alias Embedding.t5Space := IsEmbedding.t5Space - --- see Note [lower instance priority] -/-- A `T₅` space is a `T₄` space. -/ -instance (priority := 100) T5Space.toT4Space [T5Space X] : T4Space X where - -- follows from type-class inference - -/-- A subspace of a T₅ space is a T₅ space. -/ -instance [T5Space X] {p : X → Prop} : T5Space { x // p x } := - IsEmbedding.subtypeVal.t5Space - -instance ULift.instT5Space [T5Space X] : T5Space (ULift X) := - IsEmbedding.uliftDown.t5Space - -open SeparationQuotient - -/-- The `SeparationQuotient` of a completely normal R₀ space is a T₅ space. -/ -instance [CompletelyNormalSpace X] [R0Space X] : T5Space (SeparationQuotient X) where - t1 := by - rwa [((t1Space_TFAE (SeparationQuotient X)).out 1 0 :), SeparationQuotient.t1Space_iff] - completely_normal s t hd₁ hd₂ := by - rw [← disjoint_comap_iff surjective_mk, comap_mk_nhdsSet, comap_mk_nhdsSet] - apply completely_normal <;> rw [← preimage_mk_closure] - exacts [hd₁.preimage mk, hd₂.preimage mk] - -end CompletelyNormal - -/-- In a compact T₂ space, the connected component of a point equals the intersection of all -its clopen neighbourhoods. -/ -theorem connectedComponent_eq_iInter_isClopen [T2Space X] [CompactSpace X] (x : X) : - connectedComponent x = ⋂ s : { s : Set X // IsClopen s ∧ x ∈ s }, s := by - apply Subset.antisymm connectedComponent_subset_iInter_isClopen - -- Reduce to showing that the clopen intersection is connected. - refine IsPreconnected.subset_connectedComponent ?_ (mem_iInter.2 fun s => s.2.2) - -- We do this by showing that any disjoint cover by two closed sets implies - -- that one of these closed sets must contain our whole thing. - -- To reduce to the case where the cover is disjoint on all of `X` we need that `s` is closed - have hs : @IsClosed X _ (⋂ s : { s : Set X // IsClopen s ∧ x ∈ s }, s) := - isClosed_iInter fun s => s.2.1.1 - rw [isPreconnected_iff_subset_of_fully_disjoint_closed hs] - intro a b ha hb hab ab_disj - -- Since our space is normal, we get two larger disjoint open sets containing the disjoint - -- closed sets. If we can show that our intersection is a subset of any of these we can then - -- "descend" this to show that it is a subset of either a or b. - rcases normal_separation ha hb ab_disj with ⟨u, v, hu, hv, hau, hbv, huv⟩ - obtain ⟨s, H⟩ : ∃ s : Set X, IsClopen s ∧ x ∈ s ∧ s ⊆ u ∪ v := by - /- Now we find a clopen set `s` around `x`, contained in `u ∪ v`. We utilize the fact that - `X \ u ∪ v` will be compact, so there must be some finite intersection of clopen neighbourhoods - of `X` disjoint to it, but a finite intersection of clopen sets is clopen, - so we let this be our `s`. -/ - have H1 := (hu.union hv).isClosed_compl.isCompact.inter_iInter_nonempty - (fun s : { s : Set X // IsClopen s ∧ x ∈ s } => s) fun s => s.2.1.1 - rw [← not_disjoint_iff_nonempty_inter, imp_not_comm, not_forall] at H1 - cases' H1 (disjoint_compl_left_iff_subset.2 <| hab.trans <| union_subset_union hau hbv) - with si H2 - refine ⟨⋂ U ∈ si, Subtype.val U, ?_, ?_, ?_⟩ - · exact isClopen_biInter_finset fun s _ => s.2.1 - · exact mem_iInter₂.2 fun s _ => s.2.2 - · rwa [← disjoint_compl_left_iff_subset, disjoint_iff_inter_eq_empty, - ← not_nonempty_iff_eq_empty] - -- So, we get a disjoint decomposition `s = s ∩ u ∪ s ∩ v` of clopen sets. The intersection of all - -- clopen neighbourhoods will then lie in whichever of u or v x lies in and hence will be a subset - -- of either a or b. - · have H1 := isClopen_inter_of_disjoint_cover_clopen H.1 H.2.2 hu hv huv - rw [union_comm] at H - have H2 := isClopen_inter_of_disjoint_cover_clopen H.1 H.2.2 hv hu huv.symm - by_cases hxu : x ∈ u <;> [left; right] - -- The x ∈ u case. - · suffices ⋂ s : { s : Set X // IsClopen s ∧ x ∈ s }, ↑s ⊆ u - from Disjoint.left_le_of_le_sup_right hab (huv.mono this hbv) - · apply Subset.trans _ s.inter_subset_right - exact iInter_subset (fun s : { s : Set X // IsClopen s ∧ x ∈ s } => s.1) - ⟨s ∩ u, H1, mem_inter H.2.1 hxu⟩ - -- If x ∉ u, we get x ∈ v since x ∈ u ∪ v. The rest is then like the x ∈ u case. - · have h1 : x ∈ v := - (hab.trans (union_subset_union hau hbv) (mem_iInter.2 fun i => i.2.2)).resolve_left hxu - suffices ⋂ s : { s : Set X // IsClopen s ∧ x ∈ s }, ↑s ⊆ v - from (huv.symm.mono this hau).left_le_of_le_sup_left hab - · refine Subset.trans ?_ s.inter_subset_right - exact iInter_subset (fun s : { s : Set X // IsClopen s ∧ x ∈ s } => s.1) - ⟨s ∩ v, H2, mem_inter H.2.1 h1⟩ - -section Profinite - -/-- A T1 space with a clopen basis is totally separated. -/ -theorem totallySeparatedSpace_of_t1_of_basis_clopen [T1Space X] - (h : IsTopologicalBasis { s : Set X | IsClopen s }) : TotallySeparatedSpace X := by - constructor - rintro x - y - hxy - rcases h.mem_nhds_iff.mp (isOpen_ne.mem_nhds hxy) with ⟨U, hU, hxU, hyU⟩ - exact ⟨U, Uᶜ, hU.isOpen, hU.compl.isOpen, hxU, fun h => hyU h rfl, (union_compl_self U).superset, - disjoint_compl_right⟩ - -variable [T2Space X] [CompactSpace X] - -/-- A compact Hausdorff space is totally disconnected if and only if it is totally separated, this - is also true for locally compact spaces. -/ -theorem compact_t2_tot_disc_iff_tot_sep : TotallyDisconnectedSpace X ↔ TotallySeparatedSpace X := by - refine ⟨fun h => ⟨fun x _ y _ => ?_⟩, @TotallySeparatedSpace.totallyDisconnectedSpace _ _⟩ - contrapose! - intro hyp - suffices x ∈ connectedComponent y by - simpa [totallyDisconnectedSpace_iff_connectedComponent_singleton.1 h y, mem_singleton_iff] - rw [connectedComponent_eq_iInter_isClopen, mem_iInter] - rintro ⟨w : Set X, hw : IsClopen w, hy : y ∈ w⟩ - by_contra hx - exact hyp ⟨wᶜ, w, hw.1.isOpen_compl, hw.2, hx, hy, (@isCompl_compl _ w _).symm.codisjoint.top_le, - disjoint_compl_left⟩ - -variable [TotallyDisconnectedSpace X] - -/-- A totally disconnected compact Hausdorff space is totally separated. -/ -instance (priority := 100) : TotallySeparatedSpace X := - compact_t2_tot_disc_iff_tot_sep.mp inferInstance - -theorem nhds_basis_clopen (x : X) : (𝓝 x).HasBasis (fun s : Set X => x ∈ s ∧ IsClopen s) id := - ⟨fun U => by - constructor - · have hx : connectedComponent x = {x} := - totallyDisconnectedSpace_iff_connectedComponent_singleton.mp ‹_› x - rw [connectedComponent_eq_iInter_isClopen] at hx - intro hU - let N := { s // IsClopen s ∧ x ∈ s } - rsuffices ⟨⟨s, hs, hs'⟩, hs''⟩ : ∃ s : N, s.val ⊆ U - · exact ⟨s, ⟨hs', hs⟩, hs''⟩ - haveI : Nonempty N := ⟨⟨univ, isClopen_univ, mem_univ x⟩⟩ - have hNcl : ∀ s : N, IsClosed s.val := fun s => s.property.1.1 - have hdir : Directed Superset fun s : N => s.val := by - rintro ⟨s, hs, hxs⟩ ⟨t, ht, hxt⟩ - exact ⟨⟨s ∩ t, hs.inter ht, ⟨hxs, hxt⟩⟩, inter_subset_left, inter_subset_right⟩ - have h_nhd : ∀ y ∈ ⋂ s : N, s.val, U ∈ 𝓝 y := fun y y_in => by - rw [hx, mem_singleton_iff] at y_in - rwa [y_in] - exact exists_subset_nhds_of_compactSpace hdir hNcl h_nhd - · rintro ⟨V, ⟨hxV, -, V_op⟩, hUV : V ⊆ U⟩ - rw [mem_nhds_iff] - exact ⟨V, hUV, V_op, hxV⟩⟩ - -theorem isTopologicalBasis_isClopen : IsTopologicalBasis { s : Set X | IsClopen s } := by - apply isTopologicalBasis_of_isOpen_of_nhds fun U (hU : IsClopen U) => hU.2 - intro x U hxU U_op - have : U ∈ 𝓝 x := IsOpen.mem_nhds U_op hxU - rcases (nhds_basis_clopen x).mem_iff.mp this with ⟨V, ⟨hxV, hV⟩, hVU : V ⊆ U⟩ - use V - tauto - -/-- Every member of an open set in a compact Hausdorff totally disconnected space - is contained in a clopen set contained in the open set. -/ -theorem compact_exists_isClopen_in_isOpen {x : X} {U : Set X} (is_open : IsOpen U) (memU : x ∈ U) : - ∃ V : Set X, IsClopen V ∧ x ∈ V ∧ V ⊆ U := - isTopologicalBasis_isClopen.mem_nhds_iff.1 (is_open.mem_nhds memU) - -end Profinite - -section LocallyCompact - -variable {H : Type*} [TopologicalSpace H] [LocallyCompactSpace H] [T2Space H] - -/-- A locally compact Hausdorff totally disconnected space has a basis with clopen elements. -/ -theorem loc_compact_Haus_tot_disc_of_zero_dim [TotallyDisconnectedSpace H] : - IsTopologicalBasis { s : Set H | IsClopen s } := by - refine isTopologicalBasis_of_isOpen_of_nhds (fun u hu => hu.2) fun x U memU hU => ?_ - obtain ⟨s, comp, xs, sU⟩ := exists_compact_subset hU memU - let u : Set s := ((↑) : s → H) ⁻¹' interior s - have u_open_in_s : IsOpen u := isOpen_interior.preimage continuous_subtype_val - lift x to s using interior_subset xs - haveI : CompactSpace s := isCompact_iff_compactSpace.1 comp - obtain ⟨V : Set s, VisClopen, Vx, V_sub⟩ := compact_exists_isClopen_in_isOpen u_open_in_s xs - have VisClopen' : IsClopen (((↑) : s → H) '' V) := by - refine ⟨comp.isClosed.isClosedEmbedding_subtypeVal.isClosed_iff_image_isClosed.1 VisClopen.1, - ?_⟩ - let v : Set u := ((↑) : u → s) ⁻¹' V - have : ((↑) : u → H) = ((↑) : s → H) ∘ ((↑) : u → s) := rfl - have f0 : IsEmbedding ((↑) : u → H) := IsEmbedding.subtypeVal.comp IsEmbedding.subtypeVal - have f1 : IsOpenEmbedding ((↑) : u → H) := by - refine ⟨f0, ?_⟩ - · have : Set.range ((↑) : u → H) = interior s := by - rw [this, Set.range_comp, Subtype.range_coe, Subtype.image_preimage_coe] - apply Set.inter_eq_self_of_subset_right interior_subset - rw [this] - apply isOpen_interior - have f2 : IsOpen v := VisClopen.2.preimage continuous_subtype_val - have f3 : ((↑) : s → H) '' V = ((↑) : u → H) '' v := by - rw [this, image_comp, Subtype.image_preimage_coe, inter_eq_self_of_subset_right V_sub] - rw [f3] - apply f1.isOpenMap v f2 - use (↑) '' V, VisClopen', by simp [Vx], Subset.trans (by simp) sU - -/-- A locally compact Hausdorff space is totally disconnected - if and only if it is totally separated. -/ -theorem loc_compact_t2_tot_disc_iff_tot_sep : - TotallyDisconnectedSpace H ↔ TotallySeparatedSpace H := by - constructor - · intro h - exact totallySeparatedSpace_of_t1_of_basis_clopen loc_compact_Haus_tot_disc_of_zero_dim - apply TotallySeparatedSpace.totallyDisconnectedSpace - -end LocallyCompact - -/-- `ConnectedComponents X` is Hausdorff when `X` is Hausdorff and compact -/ -instance ConnectedComponents.t2 [T2Space X] [CompactSpace X] : T2Space (ConnectedComponents X) := by - -- Proof follows that of: https://stacks.math.columbia.edu/tag/0900 - -- Fix 2 distinct connected components, with points a and b - refine ⟨ConnectedComponents.surjective_coe.forall₂.2 fun a b ne => ?_⟩ - rw [ConnectedComponents.coe_ne_coe] at ne - have h := connectedComponent_disjoint ne - -- write ↑b as the intersection of all clopen subsets containing it - rw [connectedComponent_eq_iInter_isClopen b, disjoint_iff_inter_eq_empty] at h - -- Now we show that this can be reduced to some clopen containing `↑b` being disjoint to `↑a` - obtain ⟨U, V, hU, ha, hb, rfl⟩ : ∃ (U : Set X) (V : Set (ConnectedComponents X)), - IsClopen U ∧ connectedComponent a ∩ U = ∅ ∧ connectedComponent b ⊆ U ∧ (↑) ⁻¹' V = U := by - have h := - (isClosed_connectedComponent (α := X)).isCompact.elim_finite_subfamily_closed - _ (fun s : { s : Set X // IsClopen s ∧ b ∈ s } => s.2.1.1) h - cases' h with fin_a ha - -- This clopen and its complement will separate the connected components of `a` and `b` - set U : Set X := ⋂ (i : { s // IsClopen s ∧ b ∈ s }) (_ : i ∈ fin_a), i - have hU : IsClopen U := isClopen_biInter_finset fun i _ => i.2.1 - exact ⟨U, (↑) '' U, hU, ha, subset_iInter₂ fun s _ => s.2.1.connectedComponent_subset s.2.2, - (connectedComponents_preimage_image U).symm ▸ hU.biUnion_connectedComponent_eq⟩ - rw [ConnectedComponents.isQuotientMap_coe.isClopen_preimage] at hU - refine ⟨Vᶜ, V, hU.compl.isOpen, hU.isOpen, ?_, hb mem_connectedComponent, disjoint_compl_left⟩ - exact fun h => flip Set.Nonempty.ne_empty ha ⟨a, mem_connectedComponent, h⟩ - -set_option linter.style.longFile 2700 diff --git a/Mathlib/Topology/CompletelyRegular.lean b/Mathlib/Topology/Separation/CompletelyRegular.lean similarity index 100% rename from Mathlib/Topology/CompletelyRegular.lean rename to Mathlib/Topology/Separation/CompletelyRegular.lean diff --git a/Mathlib/Topology/CountableSeparatingOn.lean b/Mathlib/Topology/Separation/CountableSeparatingOn.lean similarity index 100% rename from Mathlib/Topology/CountableSeparatingOn.lean rename to Mathlib/Topology/Separation/CountableSeparatingOn.lean diff --git a/Mathlib/Topology/Separation/GDelta.lean b/Mathlib/Topology/Separation/GDelta.lean index 739f641064df5..abc821f17ade6 100644 --- a/Mathlib/Topology/Separation/GDelta.lean +++ b/Mathlib/Topology/Separation/GDelta.lean @@ -7,7 +7,7 @@ import Mathlib.Topology.Compactness.Lindelof import Mathlib.Topology.Compactness.SigmaCompact import Mathlib.Topology.Connected.TotallyDisconnected import Mathlib.Topology.Inseparable -import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.Separation.Regular import Mathlib.Topology.GDelta.Basic /-! diff --git a/Mathlib/Topology/Separation/Hausdorff.lean b/Mathlib/Topology/Separation/Hausdorff.lean new file mode 100644 index 0000000000000..17cc8d6659b95 --- /dev/null +++ b/Mathlib/Topology/Separation/Hausdorff.lean @@ -0,0 +1,666 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Mario Carneiro +-/ +import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.Compactness.SigmaCompact + +/-! +# T₂ and T₂.₅ spaces. + +This file defines the T₂ (Hausdorff) condition, which is the most commonly-used among the various +separation axioms, and the related T₂.₅ condition. + +## Main definitions + +* `T2Space`: A T₂/Hausdorff space is a space where, for every two points `x ≠ y`, + there is two disjoint open sets, one containing `x`, and the other `y`. T₂ implies T₁ and R₁. +* `T25Space`: A T₂.₅/Urysohn space is a space where, for every two points `x ≠ y`, + there is two open sets, one containing `x`, and the other `y`, whose closures are disjoint. + T₂.₅ implies T₂. + +See `Mathlib.Topology.Separation.Regular` for regular, T₃, etc spaces; and +`Mathlib.Topology.Separation.GDelta` for the definitions of `PerfectlyNormalSpace` and `T6Space`. + +Note that `mathlib` adopts the modern convention that `m ≤ n` if and only if `T_m → T_n`, but +occasionally the literature swaps definitions for e.g. T₃ and regular. + +## Main results + +### T₂ spaces + +* `t2_iff_nhds`: A space is T₂ iff the neighbourhoods of distinct points generate the bottom filter. +* `t2_iff_isClosed_diagonal`: A space is T₂ iff the `diagonal` of `X` (that is, the set of all + points of the form `(a, a) : X × X`) is closed under the product topology. +* `separatedNhds_of_finset_finset`: Any two disjoint finsets are `SeparatedNhds`. +* Most topological constructions preserve Hausdorffness; + these results are part of the typeclass inference system (e.g. `Topology.IsEmbedding.t2Space`) +* `Set.EqOn.closure`: If two functions are equal on some set `s`, they are equal on its closure. +* `IsCompact.isClosed`: All compact sets are closed. +* `WeaklyLocallyCompactSpace.locallyCompactSpace`: If a topological space is both + weakly locally compact (i.e., each point has a compact neighbourhood) + and is T₂, then it is locally compact. +* `totallySeparatedSpace_of_t1_of_basis_clopen`: If `X` has a clopen basis, then + it is a `TotallySeparatedSpace`. +* `loc_compact_t2_tot_disc_iff_tot_sep`: A locally compact T₂ space is totally disconnected iff + it is totally separated. +* `t2Quotient`: the largest T2 quotient of a given topological space. + +If the space is also compact: + +* `normalOfCompactT2`: A compact T₂ space is a `NormalSpace`. +* `connectedComponent_eq_iInter_isClopen`: The connected component of a point + is the intersection of all its clopen neighbourhoods. +* `compact_t2_tot_disc_iff_tot_sep`: Being a `TotallyDisconnectedSpace` + is equivalent to being a `TotallySeparatedSpace`. +* `ConnectedComponents.t2`: `ConnectedComponents X` is T₂ for `X` T₂ and compact. + +## References + +* +* [Willard's *General Topology*][zbMATH02107988] + +-/ + +open Function Set Filter Topology TopologicalSpace + +universe u v + +variable {X : Type*} {Y : Type*} [TopologicalSpace X] + +section Separation + +/-- A T₂ space, also known as a Hausdorff space, is one in which for every + `x ≠ y` there exists disjoint open sets around `x` and `y`. This is + the most widely used of the separation axioms. -/ +@[mk_iff] +class T2Space (X : Type u) [TopologicalSpace X] : Prop where + /-- Every two points in a Hausdorff space admit disjoint open neighbourhoods. -/ + t2 : Pairwise fun x y => ∃ u v : Set X, IsOpen u ∧ IsOpen v ∧ x ∈ u ∧ y ∈ v ∧ Disjoint u v + +/-- Two different points can be separated by open sets. -/ +theorem t2_separation [T2Space X] {x y : X} (h : x ≠ y) : + ∃ u v : Set X, IsOpen u ∧ IsOpen v ∧ x ∈ u ∧ y ∈ v ∧ Disjoint u v := + T2Space.t2 h + +-- todo: use this as a definition? +theorem t2Space_iff_disjoint_nhds : T2Space X ↔ Pairwise fun x y : X => Disjoint (𝓝 x) (𝓝 y) := by + refine (t2Space_iff X).trans (forall₃_congr fun x y _ => ?_) + simp only [(nhds_basis_opens x).disjoint_iff (nhds_basis_opens y), exists_prop, ← exists_and_left, + and_assoc, and_comm, and_left_comm] + +@[simp] +theorem disjoint_nhds_nhds [T2Space X] {x y : X} : Disjoint (𝓝 x) (𝓝 y) ↔ x ≠ y := + ⟨fun hd he => by simp [he, nhds_neBot.ne] at hd, (t2Space_iff_disjoint_nhds.mp ‹_› ·)⟩ + +theorem pairwise_disjoint_nhds [T2Space X] : Pairwise (Disjoint on (𝓝 : X → Filter X)) := fun _ _ => + disjoint_nhds_nhds.2 + +protected theorem Set.pairwiseDisjoint_nhds [T2Space X] (s : Set X) : s.PairwiseDisjoint 𝓝 := + pairwise_disjoint_nhds.set_pairwise s + +/-- Points of a finite set can be separated by open sets from each other. -/ +theorem Set.Finite.t2_separation [T2Space X] {s : Set X} (hs : s.Finite) : + ∃ U : X → Set X, (∀ x, x ∈ U x ∧ IsOpen (U x)) ∧ s.PairwiseDisjoint U := + s.pairwiseDisjoint_nhds.exists_mem_filter_basis hs nhds_basis_opens + +-- see Note [lower instance priority] +instance (priority := 100) T2Space.t1Space [T2Space X] : T1Space X := + t1Space_iff_disjoint_pure_nhds.mpr fun _ _ hne => + (disjoint_nhds_nhds.2 hne).mono_left <| pure_le_nhds _ + +-- see Note [lower instance priority] +instance (priority := 100) T2Space.r1Space [T2Space X] : R1Space X := + ⟨fun x y ↦ (eq_or_ne x y).imp specializes_of_eq disjoint_nhds_nhds.2⟩ + +theorem SeparationQuotient.t2Space_iff : T2Space (SeparationQuotient X) ↔ R1Space X := by + simp only [t2Space_iff_disjoint_nhds, Pairwise, surjective_mk.forall₂, ne_eq, mk_eq_mk, + r1Space_iff_inseparable_or_disjoint_nhds, ← disjoint_comap_iff surjective_mk, comap_mk_nhds_mk, + ← or_iff_not_imp_left] + +instance SeparationQuotient.t2Space [R1Space X] : T2Space (SeparationQuotient X) := + t2Space_iff.2 ‹_› + +instance (priority := 80) [R1Space X] [T0Space X] : T2Space X := + t2Space_iff_disjoint_nhds.2 fun _x _y hne ↦ disjoint_nhds_nhds_iff_not_inseparable.2 fun hxy ↦ + hne hxy.eq + +theorem R1Space.t2Space_iff_t0Space [R1Space X] : T2Space X ↔ T0Space X := by + constructor <;> intro <;> infer_instance + +/-- A space is T₂ iff the neighbourhoods of distinct points generate the bottom filter. -/ +theorem t2_iff_nhds : T2Space X ↔ ∀ {x y : X}, NeBot (𝓝 x ⊓ 𝓝 y) → x = y := by + simp only [t2Space_iff_disjoint_nhds, disjoint_iff, neBot_iff, Ne, not_imp_comm, Pairwise] + +theorem eq_of_nhds_neBot [T2Space X] {x y : X} (h : NeBot (𝓝 x ⊓ 𝓝 y)) : x = y := + t2_iff_nhds.mp ‹_› h + +theorem t2Space_iff_nhds : + T2Space X ↔ Pairwise fun x y : X => ∃ U ∈ 𝓝 x, ∃ V ∈ 𝓝 y, Disjoint U V := by + simp only [t2Space_iff_disjoint_nhds, Filter.disjoint_iff, Pairwise] + +theorem t2_separation_nhds [T2Space X] {x y : X} (h : x ≠ y) : + ∃ u v, u ∈ 𝓝 x ∧ v ∈ 𝓝 y ∧ Disjoint u v := + let ⟨u, v, open_u, open_v, x_in, y_in, huv⟩ := t2_separation h + ⟨u, v, open_u.mem_nhds x_in, open_v.mem_nhds y_in, huv⟩ + +theorem t2_separation_compact_nhds [LocallyCompactSpace X] [T2Space X] {x y : X} (h : x ≠ y) : + ∃ u v, u ∈ 𝓝 x ∧ v ∈ 𝓝 y ∧ IsCompact u ∧ IsCompact v ∧ Disjoint u v := by + simpa only [exists_prop, ← exists_and_left, and_comm, and_assoc, and_left_comm] using + ((compact_basis_nhds x).disjoint_iff (compact_basis_nhds y)).1 (disjoint_nhds_nhds.2 h) + +theorem t2_iff_ultrafilter : + T2Space X ↔ ∀ {x y : X} (f : Ultrafilter X), ↑f ≤ 𝓝 x → ↑f ≤ 𝓝 y → x = y := + t2_iff_nhds.trans <| by simp only [← exists_ultrafilter_iff, and_imp, le_inf_iff, exists_imp] + +theorem t2_iff_isClosed_diagonal : T2Space X ↔ IsClosed (diagonal X) := by + simp only [t2Space_iff_disjoint_nhds, ← isOpen_compl_iff, isOpen_iff_mem_nhds, Prod.forall, + nhds_prod_eq, compl_diagonal_mem_prod, mem_compl_iff, mem_diagonal_iff, Pairwise] + +theorem isClosed_diagonal [T2Space X] : IsClosed (diagonal X) := + t2_iff_isClosed_diagonal.mp ‹_› + +-- Porting note: 2 lemmas moved below + +theorem tendsto_nhds_unique [T2Space X] {f : Y → X} {l : Filter Y} {a b : X} [NeBot l] + (ha : Tendsto f l (𝓝 a)) (hb : Tendsto f l (𝓝 b)) : a = b := + (tendsto_nhds_unique_inseparable ha hb).eq + +theorem tendsto_nhds_unique' [T2Space X] {f : Y → X} {l : Filter Y} {a b : X} (_ : NeBot l) + (ha : Tendsto f l (𝓝 a)) (hb : Tendsto f l (𝓝 b)) : a = b := + tendsto_nhds_unique ha hb + +theorem tendsto_nhds_unique_of_eventuallyEq [T2Space X] {f g : Y → X} {l : Filter Y} {a b : X} + [NeBot l] (ha : Tendsto f l (𝓝 a)) (hb : Tendsto g l (𝓝 b)) (hfg : f =ᶠ[l] g) : a = b := + tendsto_nhds_unique (ha.congr' hfg) hb + +theorem tendsto_nhds_unique_of_frequently_eq [T2Space X] {f g : Y → X} {l : Filter Y} {a b : X} + (ha : Tendsto f l (𝓝 a)) (hb : Tendsto g l (𝓝 b)) (hfg : ∃ᶠ x in l, f x = g x) : a = b := + have : ∃ᶠ z : X × X in 𝓝 (a, b), z.1 = z.2 := (ha.prod_mk_nhds hb).frequently hfg + not_not.1 fun hne => this (isClosed_diagonal.isOpen_compl.mem_nhds hne) + +/-- If `s` and `t` are compact sets in a T₂ space, then the set neighborhoods filter of `s ∩ t` +is the infimum of set neighborhoods filters for `s` and `t`. + +For general sets, only the `≤` inequality holds, see `nhdsSet_inter_le`. -/ +theorem IsCompact.nhdsSet_inter_eq [T2Space X] {s t : Set X} (hs : IsCompact s) (ht : IsCompact t) : + 𝓝ˢ (s ∩ t) = 𝓝ˢ s ⊓ 𝓝ˢ t := by + refine le_antisymm (nhdsSet_inter_le _ _) ?_ + simp_rw [hs.nhdsSet_inf_eq_biSup, ht.inf_nhdsSet_eq_biSup, nhdsSet, sSup_image] + refine iSup₂_le fun x hxs ↦ iSup₂_le fun y hyt ↦ ?_ + rcases eq_or_ne x y with (rfl|hne) + · exact le_iSup₂_of_le x ⟨hxs, hyt⟩ (inf_idem _).le + · exact (disjoint_nhds_nhds.mpr hne).eq_bot ▸ bot_le + +/-- In a `T2Space X`, for a compact set `t` and a point `x` outside `t`, there are open sets `U`, +`V` that separate `t` and `x`.-/ +lemma IsCompact.separation_of_not_mem {X : Type u_1} [TopologicalSpace X] [T2Space X] {x : X} + {t : Set X} (H1 : IsCompact t) (H2 : x ∉ t) : + ∃ (U : Set X), ∃ (V : Set X), IsOpen U ∧ IsOpen V ∧ t ⊆ U ∧ x ∈ V ∧ Disjoint U V := by + simpa [SeparatedNhds] using SeparatedNhds.of_isCompact_isCompact_isClosed H1 isCompact_singleton + isClosed_singleton <| disjoint_singleton_right.mpr H2 + +/-- In a `T2Space X`, for a compact set `t` and a point `x` outside `t`, `𝓝ˢ t` and `𝓝 x` are +disjoint. -/ +lemma IsCompact.disjoint_nhdsSet_nhds {X : Type u_1} [TopologicalSpace X] [T2Space X] {x : X} + {t : Set X} (H1 : IsCompact t) (H2 : x ∉ t) : + Disjoint (𝓝ˢ t) (𝓝 x) := by + simpa using SeparatedNhds.disjoint_nhdsSet <| .of_isCompact_isCompact_isClosed H1 + isCompact_singleton isClosed_singleton <| disjoint_singleton_right.mpr H2 + +/-- If a function `f` is + +- injective on a compact set `s`; +- continuous at every point of this set; +- injective on a neighborhood of each point of this set, + +then it is injective on a neighborhood of this set. -/ +theorem Set.InjOn.exists_mem_nhdsSet {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] + [T2Space Y] {f : X → Y} {s : Set X} (inj : InjOn f s) (sc : IsCompact s) + (fc : ∀ x ∈ s, ContinuousAt f x) (loc : ∀ x ∈ s, ∃ u ∈ 𝓝 x, InjOn f u) : + ∃ t ∈ 𝓝ˢ s, InjOn f t := by + have : ∀ x ∈ s ×ˢ s, ∀ᶠ y in 𝓝 x, f y.1 = f y.2 → y.1 = y.2 := fun (x, y) ⟨hx, hy⟩ ↦ by + rcases eq_or_ne x y with rfl | hne + · rcases loc x hx with ⟨u, hu, hf⟩ + exact Filter.mem_of_superset (prod_mem_nhds hu hu) <| forall_prod_set.2 hf + · suffices ∀ᶠ z in 𝓝 (x, y), f z.1 ≠ f z.2 from this.mono fun _ hne h ↦ absurd h hne + refine (fc x hx).prodMap' (fc y hy) <| isClosed_diagonal.isOpen_compl.mem_nhds ?_ + exact inj.ne hx hy hne + rw [← eventually_nhdsSet_iff_forall, sc.nhdsSet_prod_eq sc] at this + exact eventually_prod_self_iff.1 this + +/-- If a function `f` is + +- injective on a compact set `s`; +- continuous at every point of this set; +- injective on a neighborhood of each point of this set, + +then it is injective on an open neighborhood of this set. -/ +theorem Set.InjOn.exists_isOpen_superset {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] + [T2Space Y] {f : X → Y} {s : Set X} (inj : InjOn f s) (sc : IsCompact s) + (fc : ∀ x ∈ s, ContinuousAt f x) (loc : ∀ x ∈ s, ∃ u ∈ 𝓝 x, InjOn f u) : + ∃ t, IsOpen t ∧ s ⊆ t ∧ InjOn f t := + let ⟨_t, hst, ht⟩ := inj.exists_mem_nhdsSet sc fc loc + let ⟨u, huo, hsu, hut⟩ := mem_nhdsSet_iff_exists.1 hst + ⟨u, huo, hsu, ht.mono hut⟩ + +section limUnder + +variable [T2Space X] {f : Filter X} + +/-! +### Properties of `lim` and `limUnder` + +In this section we use explicit `Nonempty X` instances for `lim` and `limUnder`. This way the lemmas +are useful without a `Nonempty X` instance. +-/ + + +theorem lim_eq {x : X} [NeBot f] (h : f ≤ 𝓝 x) : @lim _ _ ⟨x⟩ f = x := + tendsto_nhds_unique (le_nhds_lim ⟨x, h⟩) h + +theorem lim_eq_iff [NeBot f] (h : ∃ x : X, f ≤ 𝓝 x) {x} : @lim _ _ ⟨x⟩ f = x ↔ f ≤ 𝓝 x := + ⟨fun c => c ▸ le_nhds_lim h, lim_eq⟩ + +theorem Ultrafilter.lim_eq_iff_le_nhds [CompactSpace X] {x : X} {F : Ultrafilter X} : + F.lim = x ↔ ↑F ≤ 𝓝 x := + ⟨fun h => h ▸ F.le_nhds_lim, lim_eq⟩ + +theorem isOpen_iff_ultrafilter' [CompactSpace X] (U : Set X) : + IsOpen U ↔ ∀ F : Ultrafilter X, F.lim ∈ U → U ∈ F.1 := by + rw [isOpen_iff_ultrafilter] + refine ⟨fun h F hF => h F.lim hF F F.le_nhds_lim, ?_⟩ + intro cond x hx f h + rw [← Ultrafilter.lim_eq_iff_le_nhds.2 h] at hx + exact cond _ hx + +theorem Filter.Tendsto.limUnder_eq {x : X} {f : Filter Y} [NeBot f] {g : Y → X} + (h : Tendsto g f (𝓝 x)) : @limUnder _ _ _ ⟨x⟩ f g = x := + lim_eq h + +theorem Filter.limUnder_eq_iff {f : Filter Y} [NeBot f] {g : Y → X} (h : ∃ x, Tendsto g f (𝓝 x)) + {x} : @limUnder _ _ _ ⟨x⟩ f g = x ↔ Tendsto g f (𝓝 x) := + ⟨fun c => c ▸ tendsto_nhds_limUnder h, Filter.Tendsto.limUnder_eq⟩ + +theorem Continuous.limUnder_eq [TopologicalSpace Y] {f : Y → X} (h : Continuous f) (y : Y) : + @limUnder _ _ _ ⟨f y⟩ (𝓝 y) f = f y := + (h.tendsto y).limUnder_eq + +@[simp] +theorem lim_nhds (x : X) : @lim _ _ ⟨x⟩ (𝓝 x) = x := + lim_eq le_rfl + +@[simp] +theorem limUnder_nhds_id (x : X) : @limUnder _ _ _ ⟨x⟩ (𝓝 x) id = x := + lim_nhds x + +@[simp] +theorem lim_nhdsWithin {x : X} {s : Set X} (h : x ∈ closure s) : @lim _ _ ⟨x⟩ (𝓝[s] x) = x := + haveI : NeBot (𝓝[s] x) := mem_closure_iff_clusterPt.1 h + lim_eq inf_le_left + +@[simp] +theorem limUnder_nhdsWithin_id {x : X} {s : Set X} (h : x ∈ closure s) : + @limUnder _ _ _ ⟨x⟩ (𝓝[s] x) id = x := + lim_nhdsWithin h + +end limUnder + +/-! +### `T2Space` constructions + +We use two lemmas to prove that various standard constructions generate Hausdorff spaces from +Hausdorff spaces: + +* `separated_by_continuous` says that two points `x y : X` can be separated by open neighborhoods + provided that there exists a continuous map `f : X → Y` with a Hausdorff codomain such that + `f x ≠ f y`. We use this lemma to prove that topological spaces defined using `induced` are + Hausdorff spaces. + +* `separated_by_isOpenEmbedding` says that for an open embedding `f : X → Y` of a Hausdorff space + `X`, the images of two distinct points `x y : X`, `x ≠ y` can be separated by open neighborhoods. + We use this lemma to prove that topological spaces defined using `coinduced` are Hausdorff spaces. +-/ + +-- see Note [lower instance priority] +instance (priority := 100) DiscreteTopology.toT2Space + [DiscreteTopology X] : T2Space X := + ⟨fun x y h => ⟨{x}, {y}, isOpen_discrete _, isOpen_discrete _, rfl, rfl, disjoint_singleton.2 h⟩⟩ + +theorem separated_by_continuous [TopologicalSpace Y] [T2Space Y] + {f : X → Y} (hf : Continuous f) {x y : X} (h : f x ≠ f y) : + ∃ u v : Set X, IsOpen u ∧ IsOpen v ∧ x ∈ u ∧ y ∈ v ∧ Disjoint u v := + let ⟨u, v, uo, vo, xu, yv, uv⟩ := t2_separation h + ⟨f ⁻¹' u, f ⁻¹' v, uo.preimage hf, vo.preimage hf, xu, yv, uv.preimage _⟩ + +theorem separated_by_isOpenEmbedding [TopologicalSpace Y] [T2Space X] + {f : X → Y} (hf : IsOpenEmbedding f) {x y : X} (h : x ≠ y) : + ∃ u v : Set Y, IsOpen u ∧ IsOpen v ∧ f x ∈ u ∧ f y ∈ v ∧ Disjoint u v := + let ⟨u, v, uo, vo, xu, yv, uv⟩ := t2_separation h + ⟨f '' u, f '' v, hf.isOpenMap _ uo, hf.isOpenMap _ vo, mem_image_of_mem _ xu, + mem_image_of_mem _ yv, disjoint_image_of_injective hf.injective uv⟩ + +@[deprecated (since := "2024-10-18")] +alias separated_by_openEmbedding := separated_by_isOpenEmbedding + +instance {p : X → Prop} [T2Space X] : T2Space (Subtype p) := inferInstance + +instance Prod.t2Space [T2Space X] [TopologicalSpace Y] [T2Space Y] : T2Space (X × Y) := + inferInstance + +/-- If the codomain of an injective continuous function is a Hausdorff space, then so is its +domain. -/ +theorem T2Space.of_injective_continuous [TopologicalSpace Y] [T2Space Y] {f : X → Y} + (hinj : Injective f) (hc : Continuous f) : T2Space X := + ⟨fun _ _ h => separated_by_continuous hc (hinj.ne h)⟩ + +/-- If the codomain of a topological embedding is a Hausdorff space, then so is its domain. +See also `T2Space.of_continuous_injective`. -/ +theorem Topology.IsEmbedding.t2Space [TopologicalSpace Y] [T2Space Y] {f : X → Y} + (hf : IsEmbedding f) : T2Space X := + .of_injective_continuous hf.injective hf.continuous + +@[deprecated (since := "2024-10-26")] +alias Embedding.t2Space := IsEmbedding.t2Space + +instance ULift.instT2Space [T2Space X] : T2Space (ULift X) := + IsEmbedding.uliftDown.t2Space + +instance [T2Space X] [TopologicalSpace Y] [T2Space Y] : + T2Space (X ⊕ Y) := by + constructor + rintro (x | x) (y | y) h + · exact separated_by_isOpenEmbedding .inl <| ne_of_apply_ne _ h + · exact separated_by_continuous continuous_isLeft <| by simp + · exact separated_by_continuous continuous_isLeft <| by simp + · exact separated_by_isOpenEmbedding .inr <| ne_of_apply_ne _ h + +instance Pi.t2Space {Y : X → Type v} [∀ a, TopologicalSpace (Y a)] + [∀ a, T2Space (Y a)] : T2Space (∀ a, Y a) := + inferInstance + +instance Sigma.t2Space {ι} {X : ι → Type*} [∀ i, TopologicalSpace (X i)] [∀ a, T2Space (X a)] : + T2Space (Σi, X i) := by + constructor + rintro ⟨i, x⟩ ⟨j, y⟩ neq + rcases eq_or_ne i j with (rfl | h) + · replace neq : x ≠ y := ne_of_apply_ne _ neq + exact separated_by_isOpenEmbedding .sigmaMk neq + · let _ := (⊥ : TopologicalSpace ι); have : DiscreteTopology ι := ⟨rfl⟩ + exact separated_by_continuous (continuous_def.2 fun u _ => isOpen_sigma_fst_preimage u) h + +section +variable (X) + +/-- The smallest equivalence relation on a topological space giving a T2 quotient. -/ +def t2Setoid : Setoid X := sInf {s | T2Space (Quotient s)} + +/-- The largest T2 quotient of a topological space. This construction is left-adjoint to the +inclusion of T2 spaces into all topological spaces. -/ +def t2Quotient := Quotient (t2Setoid X) + +namespace t2Quotient +variable {X} + +instance : TopologicalSpace (t2Quotient X) := + inferInstanceAs <| TopologicalSpace (Quotient _) + +/-- The map from a topological space to its largest T2 quotient. -/ +def mk : X → t2Quotient X := Quotient.mk (t2Setoid X) + +lemma mk_eq {x y : X} : mk x = mk y ↔ ∀ s : Setoid X, T2Space (Quotient s) → s x y := + Setoid.quotient_mk_sInf_eq + +variable (X) + +lemma surjective_mk : Surjective (mk : X → t2Quotient X) := Quotient.mk_surjective + +lemma continuous_mk : Continuous (mk : X → t2Quotient X) := + continuous_quotient_mk' + +variable {X} + +@[elab_as_elim] +protected lemma inductionOn {motive : t2Quotient X → Prop} (q : t2Quotient X) + (h : ∀ x, motive (t2Quotient.mk x)) : motive q := Quotient.inductionOn q h + +@[elab_as_elim] +protected lemma inductionOn₂ [TopologicalSpace Y] {motive : t2Quotient X → t2Quotient Y → Prop} + (q : t2Quotient X) (q' : t2Quotient Y) (h : ∀ x y, motive (mk x) (mk y)) : motive q q' := + Quotient.inductionOn₂ q q' h + +/-- The largest T2 quotient of a topological space is indeed T2. -/ +instance : T2Space (t2Quotient X) := by + rw [t2Space_iff] + rintro ⟨x⟩ ⟨y⟩ (h : ¬ t2Quotient.mk x = t2Quotient.mk y) + obtain ⟨s, hs, hsxy⟩ : ∃ s, T2Space (Quotient s) ∧ Quotient.mk s x ≠ Quotient.mk s y := by + simpa [t2Quotient.mk_eq] using h + exact separated_by_continuous (continuous_map_sInf (by exact hs)) hsxy + +lemma compatible {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] [T2Space Y] + {f : X → Y} (hf : Continuous f) : letI _ := t2Setoid X + ∀ (a b : X), a ≈ b → f a = f b := by + change t2Setoid X ≤ Setoid.ker f + exact sInf_le <| .of_injective_continuous + (Setoid.ker_lift_injective _) (hf.quotient_lift fun _ _ ↦ id) + +/-- The universal property of the largest T2 quotient of a topological space `X`: any continuous +map from `X` to a T2 space `Y` uniquely factors through `t2Quotient X`. This declaration builds the +factored map. Its continuity is `t2Quotient.continuous_lift`, the fact that it indeed factors the +original map is `t2Quotient.lift_mk` and uniquenes is `t2Quotient.unique_lift`. -/ +def lift {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] [T2Space Y] + {f : X → Y} (hf : Continuous f) : t2Quotient X → Y := + Quotient.lift f (t2Quotient.compatible hf) + +lemma continuous_lift {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] [T2Space Y] + {f : X → Y} (hf : Continuous f) : Continuous (t2Quotient.lift hf) := + continuous_coinduced_dom.mpr hf + +@[simp] +lemma lift_mk {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] [T2Space Y] + {f : X → Y} (hf : Continuous f) (x : X) : lift hf (mk x) = f x := + Quotient.lift_mk (s := t2Setoid X) f (t2Quotient.compatible hf) x + +lemma unique_lift {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] [T2Space Y] + {f : X → Y} (hf : Continuous f) {g : t2Quotient X → Y} (hfg : g ∘ mk = f) : + g = lift hf := by + apply surjective_mk X |>.right_cancellable |>.mp <| funext _ + simp [← hfg] + +end t2Quotient +end + +variable {Z : Type*} [TopologicalSpace Y] [TopologicalSpace Z] + +theorem isClosed_eq [T2Space X] {f g : Y → X} (hf : Continuous f) (hg : Continuous g) : + IsClosed { y : Y | f y = g y } := + continuous_iff_isClosed.mp (hf.prod_mk hg) _ isClosed_diagonal + +/-- If functions `f` and `g` are continuous on a closed set `s`, +then the set of points `x ∈ s` such that `f x = g x` is a closed set. -/ +protected theorem IsClosed.isClosed_eq [T2Space Y] {f g : X → Y} {s : Set X} (hs : IsClosed s) + (hf : ContinuousOn f s) (hg : ContinuousOn g s) : IsClosed {x ∈ s | f x = g x} := + (hf.prod hg).preimage_isClosed_of_isClosed hs isClosed_diagonal + +theorem isOpen_ne_fun [T2Space X] {f g : Y → X} (hf : Continuous f) (hg : Continuous g) : + IsOpen { y : Y | f y ≠ g y } := + isOpen_compl_iff.mpr <| isClosed_eq hf hg + +/-- If two continuous maps are equal on `s`, then they are equal on the closure of `s`. See also +`Set.EqOn.of_subset_closure` for a more general version. -/ +protected theorem Set.EqOn.closure [T2Space X] {s : Set Y} {f g : Y → X} (h : EqOn f g s) + (hf : Continuous f) (hg : Continuous g) : EqOn f g (closure s) := + closure_minimal h (isClosed_eq hf hg) + +/-- If two continuous functions are equal on a dense set, then they are equal. -/ +theorem Continuous.ext_on [T2Space X] {s : Set Y} (hs : Dense s) {f g : Y → X} (hf : Continuous f) + (hg : Continuous g) (h : EqOn f g s) : f = g := + funext fun x => h.closure hf hg (hs x) + +theorem eqOn_closure₂' [T2Space Z] {s : Set X} {t : Set Y} {f g : X → Y → Z} + (h : ∀ x ∈ s, ∀ y ∈ t, f x y = g x y) (hf₁ : ∀ x, Continuous (f x)) + (hf₂ : ∀ y, Continuous fun x => f x y) (hg₁ : ∀ x, Continuous (g x)) + (hg₂ : ∀ y, Continuous fun x => g x y) : ∀ x ∈ closure s, ∀ y ∈ closure t, f x y = g x y := + suffices closure s ⊆ ⋂ y ∈ closure t, { x | f x y = g x y } by simpa only [subset_def, mem_iInter] + (closure_minimal fun x hx => mem_iInter₂.2 <| Set.EqOn.closure (h x hx) (hf₁ _) (hg₁ _)) <| + isClosed_biInter fun _ _ => isClosed_eq (hf₂ _) (hg₂ _) + +theorem eqOn_closure₂ [T2Space Z] {s : Set X} {t : Set Y} {f g : X → Y → Z} + (h : ∀ x ∈ s, ∀ y ∈ t, f x y = g x y) (hf : Continuous (uncurry f)) + (hg : Continuous (uncurry g)) : ∀ x ∈ closure s, ∀ y ∈ closure t, f x y = g x y := + eqOn_closure₂' h hf.uncurry_left hf.uncurry_right hg.uncurry_left hg.uncurry_right + +/-- If `f x = g x` for all `x ∈ s` and `f`, `g` are continuous on `t`, `s ⊆ t ⊆ closure s`, then +`f x = g x` for all `x ∈ t`. See also `Set.EqOn.closure`. -/ +theorem Set.EqOn.of_subset_closure [T2Space Y] {s t : Set X} {f g : X → Y} (h : EqOn f g s) + (hf : ContinuousOn f t) (hg : ContinuousOn g t) (hst : s ⊆ t) (hts : t ⊆ closure s) : + EqOn f g t := by + intro x hx + have : (𝓝[s] x).NeBot := mem_closure_iff_clusterPt.mp (hts hx) + exact + tendsto_nhds_unique_of_eventuallyEq ((hf x hx).mono_left <| nhdsWithin_mono _ hst) + ((hg x hx).mono_left <| nhdsWithin_mono _ hst) (h.eventuallyEq_of_mem self_mem_nhdsWithin) + +theorem Function.LeftInverse.isClosed_range [T2Space X] {f : X → Y} {g : Y → X} + (h : Function.LeftInverse f g) (hf : Continuous f) (hg : Continuous g) : IsClosed (range g) := + have : EqOn (g ∘ f) id (closure <| range g) := + h.rightInvOn_range.eqOn.closure (hg.comp hf) continuous_id + isClosed_of_closure_subset fun x hx => ⟨f x, this hx⟩ + +@[deprecated (since := "2024-03-17")] +alias Function.LeftInverse.closed_range := Function.LeftInverse.isClosed_range + +theorem Function.LeftInverse.isClosedEmbedding [T2Space X] {f : X → Y} {g : Y → X} + (h : Function.LeftInverse f g) (hf : Continuous f) (hg : Continuous g) : IsClosedEmbedding g := + ⟨.of_leftInverse h hf hg, h.isClosed_range hf hg⟩ + +@[deprecated (since := "2024-10-20")] +alias Function.LeftInverse.closedEmbedding := Function.LeftInverse.isClosedEmbedding + +theorem SeparatedNhds.of_isCompact_isCompact [T2Space X] {s t : Set X} (hs : IsCompact s) + (ht : IsCompact t) (hst : Disjoint s t) : SeparatedNhds s t := by + simp only [SeparatedNhds, prod_subset_compl_diagonal_iff_disjoint.symm] at hst ⊢ + exact generalized_tube_lemma hs ht isClosed_diagonal.isOpen_compl hst + +/-- In a `T2Space X`, for disjoint closed sets `s t` such that `closure sᶜ` is compact, +there are neighbourhoods that separate `s` and `t`.-/ +lemma SeparatedNhds.of_isClosed_isCompact_closure_compl_isClosed [T2Space X] {s : Set X} + {t : Set X} (H1 : IsClosed s) (H2 : IsCompact (closure sᶜ)) (H3 : IsClosed t) + (H4 : Disjoint s t) : SeparatedNhds s t := by + -- Since `t` is a closed subset of the compact set `closure sᶜ`, it is compact. + have ht : IsCompact t := .of_isClosed_subset H2 H3 <| H4.subset_compl_left.trans subset_closure + -- we split `s` into its frontier and its interior. + rw [← diff_union_of_subset (interior_subset (s := s))] + -- since `t ⊆ sᶜ`, which is open, and `interior s` is open, we have + -- `SeparatedNhds (interior s) t`, which leaves us only with the frontier. + refine .union_left ?_ ⟨interior s, sᶜ, isOpen_interior, H1.isOpen_compl, le_rfl, + H4.subset_compl_left, disjoint_compl_right.mono_left interior_subset⟩ + -- Since the frontier of `s` is compact (as it is a subset of `closure sᶜ`), we simply apply + -- `SeparatedNhds_of_isCompact_isCompact`. + rw [← H1.frontier_eq, frontier_eq_closure_inter_closure, H1.closure_eq] + refine .of_isCompact_isCompact ?_ ht (disjoint_of_subset_left inter_subset_left H4) + exact H2.of_isClosed_subset (H1.inter isClosed_closure) inter_subset_right + +section SeparatedFinset + +theorem SeparatedNhds.of_finset_finset [T2Space X] (s t : Finset X) (h : Disjoint s t) : + SeparatedNhds (s : Set X) t := + .of_isCompact_isCompact s.finite_toSet.isCompact t.finite_toSet.isCompact <| mod_cast h + +theorem SeparatedNhds.of_singleton_finset [T2Space X] {x : X} {s : Finset X} (h : x ∉ s) : + SeparatedNhds ({x} : Set X) s := + mod_cast .of_finset_finset {x} s (Finset.disjoint_singleton_left.mpr h) + +end SeparatedFinset + +/-- In a `T2Space`, every compact set is closed. -/ +theorem IsCompact.isClosed [T2Space X] {s : Set X} (hs : IsCompact s) : IsClosed s := + isOpen_compl_iff.1 <| isOpen_iff_forall_mem_open.mpr fun x hx => + let ⟨u, v, _, vo, su, xv, uv⟩ := + SeparatedNhds.of_isCompact_isCompact hs isCompact_singleton (disjoint_singleton_right.2 hx) + ⟨v, (uv.mono_left <| show s ≤ u from su).subset_compl_left, vo, by simpa using xv⟩ + +theorem IsCompact.preimage_continuous [CompactSpace X] [T2Space Y] {f : X → Y} {s : Set Y} + (hs : IsCompact s) (hf : Continuous f) : IsCompact (f ⁻¹' s) := + (hs.isClosed.preimage hf).isCompact + +lemma Pi.isCompact_iff {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] + [∀ i, T2Space (π i)] {s : Set (Π i, π i)} : + IsCompact s ↔ IsClosed s ∧ ∀ i, IsCompact (eval i '' s) := by + constructor <;> intro H + · exact ⟨H.isClosed, fun i ↦ H.image <| continuous_apply i⟩ + · exact IsCompact.of_isClosed_subset (isCompact_univ_pi H.2) H.1 (subset_pi_eval_image univ s) + +lemma Pi.isCompact_closure_iff {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)] + [∀ i, T2Space (π i)] {s : Set (Π i, π i)} : + IsCompact (closure s) ↔ ∀ i, IsCompact (closure <| eval i '' s) := by + simp_rw [← exists_isCompact_superset_iff, Pi.exists_compact_superset_iff, image_subset_iff] + +/-- If `V : ι → Set X` is a decreasing family of compact sets then any neighborhood of +`⋂ i, V i` contains some `V i`. This is a version of `exists_subset_nhds_of_isCompact'` where we +don't need to assume each `V i` closed because it follows from compactness since `X` is +assumed to be Hausdorff. -/ +theorem exists_subset_nhds_of_isCompact [T2Space X] {ι : Type*} [Nonempty ι] {V : ι → Set X} + (hV : Directed (· ⊇ ·) V) (hV_cpct : ∀ i, IsCompact (V i)) {U : Set X} + (hU : ∀ x ∈ ⋂ i, V i, U ∈ 𝓝 x) : ∃ i, V i ⊆ U := + exists_subset_nhds_of_isCompact' hV hV_cpct (fun i => (hV_cpct i).isClosed) hU + +theorem CompactExhaustion.isClosed [T2Space X] (K : CompactExhaustion X) (n : ℕ) : IsClosed (K n) := + (K.isCompact n).isClosed + +theorem IsCompact.inter [T2Space X] {s t : Set X} (hs : IsCompact s) (ht : IsCompact t) : + IsCompact (s ∩ t) := + hs.inter_right <| ht.isClosed + +theorem image_closure_of_isCompact [T2Space Y] {s : Set X} (hs : IsCompact (closure s)) {f : X → Y} + (hf : ContinuousOn f (closure s)) : f '' closure s = closure (f '' s) := + Subset.antisymm hf.image_closure <| + closure_minimal (image_subset f subset_closure) (hs.image_of_continuousOn hf).isClosed + +/-- A continuous map from a compact space to a Hausdorff space is a closed map. -/ +protected theorem Continuous.isClosedMap [CompactSpace X] [T2Space Y] {f : X → Y} + (h : Continuous f) : IsClosedMap f := fun _s hs => (hs.isCompact.image h).isClosed + +/-- A continuous injective map from a compact space to a Hausdorff space is a closed embedding. -/ +theorem Continuous.isClosedEmbedding [CompactSpace X] [T2Space Y] {f : X → Y} (h : Continuous f) + (hf : Function.Injective f) : IsClosedEmbedding f := + .of_continuous_injective_isClosedMap h hf h.isClosedMap + +@[deprecated (since := "2024-10-20")] +alias Continuous.closedEmbedding := Continuous.isClosedEmbedding + +/-- A continuous surjective map from a compact space to a Hausdorff space is a quotient map. -/ +theorem IsQuotientMap.of_surjective_continuous [CompactSpace X] [T2Space Y] {f : X → Y} + (hsurj : Surjective f) (hcont : Continuous f) : IsQuotientMap f := + hcont.isClosedMap.isQuotientMap hcont hsurj + +@[deprecated (since := "2024-10-22")] +alias QuotientMap.of_surjective_continuous := IsQuotientMap.of_surjective_continuous + +theorem isPreirreducible_iff_subsingleton [T2Space X] {S : Set X} : + IsPreirreducible S ↔ S.Subsingleton := by + refine ⟨fun h x hx y hy => ?_, Set.Subsingleton.isPreirreducible⟩ + by_contra e + obtain ⟨U, V, hU, hV, hxU, hyV, h'⟩ := t2_separation e + exact ((h U V hU hV ⟨x, hx, hxU⟩ ⟨y, hy, hyV⟩).mono inter_subset_right).not_disjoint h' + +-- todo: use `alias` + `attribute [protected]` once we get `attribute [protected]` +protected lemma IsPreirreducible.subsingleton [T2Space X] {S : Set X} (h : IsPreirreducible S) : + S.Subsingleton := + isPreirreducible_iff_subsingleton.1 h + +theorem isIrreducible_iff_singleton [T2Space X] {S : Set X} : IsIrreducible S ↔ ∃ x, S = {x} := by + rw [IsIrreducible, isPreirreducible_iff_subsingleton, + exists_eq_singleton_iff_nonempty_subsingleton] + +/-- There does not exist a nontrivial preirreducible T₂ space. -/ +theorem not_preirreducible_nontrivial_t2 (X) [TopologicalSpace X] [PreirreducibleSpace X] + [Nontrivial X] [T2Space X] : False := + (PreirreducibleSpace.isPreirreducible_univ (X := X)).subsingleton.not_nontrivial nontrivial_univ + +theorem t2Space_antitone {X : Type*} : Antitone (@T2Space X) := + fun inst₁ inst₂ h_top h_t2 ↦ @T2Space.of_injective_continuous _ _ inst₁ inst₂ + h_t2 _ Function.injective_id <| continuous_id_of_le h_top + +end Separation diff --git a/Mathlib/Topology/Separation/Profinite.lean b/Mathlib/Topology/Separation/Profinite.lean new file mode 100644 index 0000000000000..e07e9b3183a3c --- /dev/null +++ b/Mathlib/Topology/Separation/Profinite.lean @@ -0,0 +1,121 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Mario Carneiro +-/ + +import Mathlib.Topology.Separation.Regular + +/-! +# Separation properties: profinite spaces +-/ + +open Function Set Filter Topology TopologicalSpace + +universe u v + +variable {X : Type*} {Y : Type*} [TopologicalSpace X] + +section Profinite + +/-- A T1 space with a clopen basis is totally separated. -/ +theorem totallySeparatedSpace_of_t1_of_basis_clopen [T1Space X] + (h : IsTopologicalBasis { s : Set X | IsClopen s }) : TotallySeparatedSpace X := by + constructor + rintro x - y - hxy + rcases h.mem_nhds_iff.mp (isOpen_ne.mem_nhds hxy) with ⟨U, hU, hxU, hyU⟩ + exact ⟨U, Uᶜ, hU.isOpen, hU.compl.isOpen, hxU, fun h => hyU h rfl, (union_compl_self U).superset, + disjoint_compl_right⟩ + +variable [T2Space X] [CompactSpace X] [TotallyDisconnectedSpace X] + +theorem nhds_basis_clopen (x : X) : (𝓝 x).HasBasis (fun s : Set X => x ∈ s ∧ IsClopen s) id := + ⟨fun U => by + constructor + · have hx : connectedComponent x = {x} := + totallyDisconnectedSpace_iff_connectedComponent_singleton.mp ‹_› x + rw [connectedComponent_eq_iInter_isClopen] at hx + intro hU + let N := { s // IsClopen s ∧ x ∈ s } + rsuffices ⟨⟨s, hs, hs'⟩, hs''⟩ : ∃ s : N, s.val ⊆ U + · exact ⟨s, ⟨hs', hs⟩, hs''⟩ + haveI : Nonempty N := ⟨⟨univ, isClopen_univ, mem_univ x⟩⟩ + have hNcl : ∀ s : N, IsClosed s.val := fun s => s.property.1.1 + have hdir : Directed Superset fun s : N => s.val := by + rintro ⟨s, hs, hxs⟩ ⟨t, ht, hxt⟩ + exact ⟨⟨s ∩ t, hs.inter ht, ⟨hxs, hxt⟩⟩, inter_subset_left, inter_subset_right⟩ + have h_nhd : ∀ y ∈ ⋂ s : N, s.val, U ∈ 𝓝 y := fun y y_in => by + rw [hx, mem_singleton_iff] at y_in + rwa [y_in] + exact exists_subset_nhds_of_compactSpace hdir hNcl h_nhd + · rintro ⟨V, ⟨hxV, -, V_op⟩, hUV : V ⊆ U⟩ + rw [mem_nhds_iff] + exact ⟨V, hUV, V_op, hxV⟩⟩ + +theorem isTopologicalBasis_isClopen : IsTopologicalBasis { s : Set X | IsClopen s } := by + apply isTopologicalBasis_of_isOpen_of_nhds fun U (hU : IsClopen U) => hU.2 + intro x U hxU U_op + have : U ∈ 𝓝 x := IsOpen.mem_nhds U_op hxU + rcases (nhds_basis_clopen x).mem_iff.mp this with ⟨V, ⟨hxV, hV⟩, hVU : V ⊆ U⟩ + use V + tauto + +/-- Every member of an open set in a compact Hausdorff totally disconnected space + is contained in a clopen set contained in the open set. -/ +theorem compact_exists_isClopen_in_isOpen {x : X} {U : Set X} (is_open : IsOpen U) (memU : x ∈ U) : + ∃ V : Set X, IsClopen V ∧ x ∈ V ∧ V ⊆ U := + isTopologicalBasis_isClopen.mem_nhds_iff.1 (is_open.mem_nhds memU) + +end Profinite + +section LocallyCompact + +variable {H : Type*} [TopologicalSpace H] [LocallyCompactSpace H] [T2Space H] + +/-- A locally compact Hausdorff totally disconnected space has a basis with clopen elements. -/ +theorem loc_compact_Haus_tot_disc_of_zero_dim [TotallyDisconnectedSpace H] : + IsTopologicalBasis { s : Set H | IsClopen s } := by + refine isTopologicalBasis_of_isOpen_of_nhds (fun u hu => hu.2) fun x U memU hU => ?_ + obtain ⟨s, comp, xs, sU⟩ := exists_compact_subset hU memU + let u : Set s := ((↑) : s → H) ⁻¹' interior s + have u_open_in_s : IsOpen u := isOpen_interior.preimage continuous_subtype_val + lift x to s using interior_subset xs + haveI : CompactSpace s := isCompact_iff_compactSpace.1 comp + obtain ⟨V : Set s, VisClopen, Vx, V_sub⟩ := compact_exists_isClopen_in_isOpen u_open_in_s xs + have VisClopen' : IsClopen (((↑) : s → H) '' V) := by + refine ⟨comp.isClosed.isClosedEmbedding_subtypeVal.isClosed_iff_image_isClosed.1 VisClopen.1, + ?_⟩ + let v : Set u := ((↑) : u → s) ⁻¹' V + have : ((↑) : u → H) = ((↑) : s → H) ∘ ((↑) : u → s) := rfl + have f0 : IsEmbedding ((↑) : u → H) := IsEmbedding.subtypeVal.comp IsEmbedding.subtypeVal + have f1 : IsOpenEmbedding ((↑) : u → H) := by + refine ⟨f0, ?_⟩ + · have : Set.range ((↑) : u → H) = interior s := by + rw [this, Set.range_comp, Subtype.range_coe, Subtype.image_preimage_coe] + apply Set.inter_eq_self_of_subset_right interior_subset + rw [this] + apply isOpen_interior + have f2 : IsOpen v := VisClopen.2.preimage continuous_subtype_val + have f3 : ((↑) : s → H) '' V = ((↑) : u → H) '' v := by + rw [this, image_comp, Subtype.image_preimage_coe, inter_eq_self_of_subset_right V_sub] + rw [f3] + apply f1.isOpenMap v f2 + use (↑) '' V, VisClopen', by simp [Vx], Subset.trans (by simp) sU + +/-- A locally compact Hausdorff space is totally disconnected + if and only if it is totally separated. -/ +theorem loc_compact_t2_tot_disc_iff_tot_sep : + TotallyDisconnectedSpace H ↔ TotallySeparatedSpace H := by + constructor + · intro h + exact totallySeparatedSpace_of_t1_of_basis_clopen loc_compact_Haus_tot_disc_of_zero_dim + apply TotallySeparatedSpace.totallyDisconnectedSpace + +@[deprecated (since := "2024-12-18")] alias compact_t2_tot_disc_iff_tot_sep := + loc_compact_t2_tot_disc_iff_tot_sep + +/-- A totally disconnected compact Hausdorff space is totally separated. -/ +instance (priority := 100) [TotallyDisconnectedSpace H] : TotallySeparatedSpace H := + loc_compact_t2_tot_disc_iff_tot_sep.mp inferInstance + +end LocallyCompact diff --git a/Mathlib/Topology/Separation/Regular.lean b/Mathlib/Topology/Separation/Regular.lean new file mode 100644 index 0000000000000..f886e9d275213 --- /dev/null +++ b/Mathlib/Topology/Separation/Regular.lean @@ -0,0 +1,671 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Mario Carneiro +-/ +import Mathlib.Topology.Separation.Hausdorff +import Mathlib.Topology.Compactness.Lindelof + +/-! +# Regular, normal, T₃, T₄ and T₅ spaces + +This file continues the study of separation properties of topological spaces, focussing +on conditions strictly stronger than T₂. + +## Main definitions + +* `RegularSpace`: A regular space is one where, given any closed `C` and `x ∉ C`, + there are disjoint open sets containing `x` and `C` respectively. Such a space is not necessarily + Hausdorff. +* `T3Space`: A T₃ space is a regular T₀ space. T₃ implies T₂.₅. +* `NormalSpace`: A normal space, is one where given two disjoint closed sets, + we can find two open sets that separate them. Such a space is not necessarily Hausdorff, even if + it is T₀. +* `T4Space`: A T₄ space is a normal T₁ space. T₄ implies T₃. +* `CompletelyNormalSpace`: A completely normal space is one in which for any two sets `s`, `t` + such that if both `closure s` is disjoint with `t`, and `s` is disjoint with `closure t`, + then there exist disjoint neighbourhoods of `s` and `t`. `Embedding.completelyNormalSpace` allows + us to conclude that this is equivalent to all subspaces being normal. Such a space is not + necessarily Hausdorff or regular, even if it is T₀. +* `T5Space`: A T₅ space is a completely normal T₁ space. T₅ implies T₄. + +See `Mathlib.Topology.Separation.GDelta` for the definitions of `PerfectlyNormalSpace` and +`T6Space`. + +Note that `mathlib` adopts the modern convention that `m ≤ n` if and only if `T_m → T_n`, but +occasionally the literature swaps definitions for e.g. T₃ and regular. + +## Main results + +### Regular spaces + +If the space is also Lindelöf: + +* `NormalSpace.of_regularSpace_lindelofSpace`: every regular Lindelöf space is normal. + +### T₃ spaces + +* `disjoint_nested_nhds`: Given two points `x ≠ y`, we can find neighbourhoods `x ∈ V₁ ⊆ U₁` and + `y ∈ V₂ ⊆ U₂`, with the `Vₖ` closed and the `Uₖ` open, such that the `Uₖ` are disjoint. + +## References + +* +* +* [Willard's *General Topology*][zbMATH02107988] + +-/ + +assert_not_exists UniformSpace + +open Function Set Filter Topology TopologicalSpace + +universe u v + +variable {X : Type*} {Y : Type*} [TopologicalSpace X] + +section RegularSpace + +/-- A topological space is called a *regular space* if for any closed set `s` and `a ∉ s`, there +exist disjoint open sets `U ⊇ s` and `V ∋ a`. We formulate this condition in terms of `Disjoint`ness +of filters `𝓝ˢ s` and `𝓝 a`. -/ +@[mk_iff] +class RegularSpace (X : Type u) [TopologicalSpace X] : Prop where + /-- If `a` is a point that does not belong to a closed set `s`, then `a` and `s` admit disjoint + neighborhoods. -/ + regular : ∀ {s : Set X} {a}, IsClosed s → a ∉ s → Disjoint (𝓝ˢ s) (𝓝 a) + +theorem regularSpace_TFAE (X : Type u) [TopologicalSpace X] : + List.TFAE [RegularSpace X, + ∀ (s : Set X) x, x ∉ closure s → Disjoint (𝓝ˢ s) (𝓝 x), + ∀ (x : X) (s : Set X), Disjoint (𝓝ˢ s) (𝓝 x) ↔ x ∉ closure s, + ∀ (x : X) (s : Set X), s ∈ 𝓝 x → ∃ t ∈ 𝓝 x, IsClosed t ∧ t ⊆ s, + ∀ x : X, (𝓝 x).lift' closure ≤ 𝓝 x, + ∀ x : X , (𝓝 x).lift' closure = 𝓝 x] := by + tfae_have 1 ↔ 5 := by + rw [regularSpace_iff, (@compl_surjective (Set X) _).forall, forall_swap] + simp only [isClosed_compl_iff, mem_compl_iff, Classical.not_not, @and_comm (_ ∈ _), + (nhds_basis_opens _).lift'_closure.le_basis_iff (nhds_basis_opens _), and_imp, + (nhds_basis_opens _).disjoint_iff_right, exists_prop, ← subset_interior_iff_mem_nhdsSet, + interior_compl, compl_subset_compl] + tfae_have 5 → 6 := fun h a => (h a).antisymm (𝓝 _).le_lift'_closure + tfae_have 6 → 4 + | H, a, s, hs => by + rw [← H] at hs + rcases (𝓝 a).basis_sets.lift'_closure.mem_iff.mp hs with ⟨U, hU, hUs⟩ + exact ⟨closure U, mem_of_superset hU subset_closure, isClosed_closure, hUs⟩ + tfae_have 4 → 2 + | H, s, a, ha => by + have ha' : sᶜ ∈ 𝓝 a := by rwa [← mem_interior_iff_mem_nhds, interior_compl] + rcases H _ _ ha' with ⟨U, hU, hUc, hUs⟩ + refine disjoint_of_disjoint_of_mem disjoint_compl_left ?_ hU + rwa [← subset_interior_iff_mem_nhdsSet, hUc.isOpen_compl.interior_eq, subset_compl_comm] + tfae_have 2 → 3 := by + refine fun H a s => ⟨fun hd has => mem_closure_iff_nhds_ne_bot.mp has ?_, H s a⟩ + exact (hd.symm.mono_right <| @principal_le_nhdsSet _ _ s).eq_bot + tfae_have 3 → 1 := fun H => ⟨fun hs ha => (H _ _).mpr <| hs.closure_eq.symm ▸ ha⟩ + tfae_finish + +theorem RegularSpace.of_lift'_closure_le (h : ∀ x : X, (𝓝 x).lift' closure ≤ 𝓝 x) : + RegularSpace X := + Iff.mpr ((regularSpace_TFAE X).out 0 4) h + +theorem RegularSpace.of_lift'_closure (h : ∀ x : X, (𝓝 x).lift' closure = 𝓝 x) : RegularSpace X := + Iff.mpr ((regularSpace_TFAE X).out 0 5) h + +@[deprecated (since := "2024-02-28")] +alias RegularSpace.ofLift'_closure := RegularSpace.of_lift'_closure + +theorem RegularSpace.of_hasBasis {ι : X → Sort*} {p : ∀ a, ι a → Prop} {s : ∀ a, ι a → Set X} + (h₁ : ∀ a, (𝓝 a).HasBasis (p a) (s a)) (h₂ : ∀ a i, p a i → IsClosed (s a i)) : + RegularSpace X := + .of_lift'_closure fun a => (h₁ a).lift'_closure_eq_self (h₂ a) + +@[deprecated (since := "2024-02-28")] +alias RegularSpace.ofBasis := RegularSpace.of_hasBasis + +theorem RegularSpace.of_exists_mem_nhds_isClosed_subset + (h : ∀ (x : X), ∀ s ∈ 𝓝 x, ∃ t ∈ 𝓝 x, IsClosed t ∧ t ⊆ s) : RegularSpace X := + Iff.mpr ((regularSpace_TFAE X).out 0 3) h + +@[deprecated (since := "2024-02-28")] +alias RegularSpace.ofExistsMemNhdsIsClosedSubset := RegularSpace.of_exists_mem_nhds_isClosed_subset + +/-- A weakly locally compact R₁ space is regular. -/ +instance (priority := 100) [WeaklyLocallyCompactSpace X] [R1Space X] : RegularSpace X := + .of_hasBasis isCompact_isClosed_basis_nhds fun _ _ ⟨_, _, h⟩ ↦ h + +section +variable [RegularSpace X] {x : X} {s : Set X} + +theorem disjoint_nhdsSet_nhds : Disjoint (𝓝ˢ s) (𝓝 x) ↔ x ∉ closure s := by + have h := (regularSpace_TFAE X).out 0 2 + exact h.mp ‹_› _ _ + +theorem disjoint_nhds_nhdsSet : Disjoint (𝓝 x) (𝓝ˢ s) ↔ x ∉ closure s := + disjoint_comm.trans disjoint_nhdsSet_nhds + +/-- A regular space is R₁. -/ +instance (priority := 100) : R1Space X where + specializes_or_disjoint_nhds _ _ := or_iff_not_imp_left.2 fun h ↦ by + rwa [← nhdsSet_singleton, disjoint_nhdsSet_nhds, ← specializes_iff_mem_closure] + +theorem exists_mem_nhds_isClosed_subset {x : X} {s : Set X} (h : s ∈ 𝓝 x) : + ∃ t ∈ 𝓝 x, IsClosed t ∧ t ⊆ s := by + have h' := (regularSpace_TFAE X).out 0 3 + exact h'.mp ‹_› _ _ h + +theorem closed_nhds_basis (x : X) : (𝓝 x).HasBasis (fun s : Set X => s ∈ 𝓝 x ∧ IsClosed s) id := + hasBasis_self.2 fun _ => exists_mem_nhds_isClosed_subset + +theorem lift'_nhds_closure (x : X) : (𝓝 x).lift' closure = 𝓝 x := + (closed_nhds_basis x).lift'_closure_eq_self fun _ => And.right + +theorem Filter.HasBasis.nhds_closure {ι : Sort*} {x : X} {p : ι → Prop} {s : ι → Set X} + (h : (𝓝 x).HasBasis p s) : (𝓝 x).HasBasis p fun i => closure (s i) := + lift'_nhds_closure x ▸ h.lift'_closure + +theorem hasBasis_nhds_closure (x : X) : (𝓝 x).HasBasis (fun s => s ∈ 𝓝 x) closure := + (𝓝 x).basis_sets.nhds_closure + +theorem hasBasis_opens_closure (x : X) : (𝓝 x).HasBasis (fun s => x ∈ s ∧ IsOpen s) closure := + (nhds_basis_opens x).nhds_closure + +theorem IsCompact.exists_isOpen_closure_subset {K U : Set X} (hK : IsCompact K) (hU : U ∈ 𝓝ˢ K) : + ∃ V, IsOpen V ∧ K ⊆ V ∧ closure V ⊆ U := by + have hd : Disjoint (𝓝ˢ K) (𝓝ˢ Uᶜ) := by + simpa [hK.disjoint_nhdsSet_left, disjoint_nhds_nhdsSet, + ← subset_interior_iff_mem_nhdsSet] using hU + rcases ((hasBasis_nhdsSet _).disjoint_iff (hasBasis_nhdsSet _)).1 hd + with ⟨V, ⟨hVo, hKV⟩, W, ⟨hW, hUW⟩, hVW⟩ + refine ⟨V, hVo, hKV, Subset.trans ?_ (compl_subset_comm.1 hUW)⟩ + exact closure_minimal hVW.subset_compl_right hW.isClosed_compl + +theorem IsCompact.lift'_closure_nhdsSet {K : Set X} (hK : IsCompact K) : + (𝓝ˢ K).lift' closure = 𝓝ˢ K := by + refine le_antisymm (fun U hU ↦ ?_) (le_lift'_closure _) + rcases hK.exists_isOpen_closure_subset hU with ⟨V, hVo, hKV, hVU⟩ + exact mem_of_superset (mem_lift' <| hVo.mem_nhdsSet.2 hKV) hVU + +theorem TopologicalSpace.IsTopologicalBasis.nhds_basis_closure {B : Set (Set X)} + (hB : IsTopologicalBasis B) (x : X) : + (𝓝 x).HasBasis (fun s : Set X => x ∈ s ∧ s ∈ B) closure := by + simpa only [and_comm] using hB.nhds_hasBasis.nhds_closure + +theorem TopologicalSpace.IsTopologicalBasis.exists_closure_subset {B : Set (Set X)} + (hB : IsTopologicalBasis B) {x : X} {s : Set X} (h : s ∈ 𝓝 x) : + ∃ t ∈ B, x ∈ t ∧ closure t ⊆ s := by + simpa only [exists_prop, and_assoc] using hB.nhds_hasBasis.nhds_closure.mem_iff.mp h + +protected theorem Topology.IsInducing.regularSpace [TopologicalSpace Y] {f : Y → X} + (hf : IsInducing f) : RegularSpace Y := + .of_hasBasis + (fun b => by rw [hf.nhds_eq_comap b]; exact (closed_nhds_basis _).comap _) + fun b s hs => by exact hs.2.preimage hf.continuous + +@[deprecated (since := "2024-10-28")] alias Inducing.regularSpace := IsInducing.regularSpace + +theorem regularSpace_induced (f : Y → X) : @RegularSpace Y (induced f ‹_›) := + letI := induced f ‹_› + (IsInducing.induced f).regularSpace + +theorem regularSpace_sInf {X} {T : Set (TopologicalSpace X)} (h : ∀ t ∈ T, @RegularSpace X t) : + @RegularSpace X (sInf T) := by + let _ := sInf T + have : ∀ a, (𝓝 a).HasBasis + (fun If : Σ I : Set T, I → Set X => + If.1.Finite ∧ ∀ i : If.1, If.2 i ∈ @nhds X i a ∧ @IsClosed X i (If.2 i)) + fun If => ⋂ i : If.1, If.snd i := fun a ↦ by + rw [nhds_sInf, ← iInf_subtype''] + exact hasBasis_iInf fun t : T => @closed_nhds_basis X t (h t t.2) a + refine .of_hasBasis this fun a If hIf => isClosed_iInter fun i => ?_ + exact (hIf.2 i).2.mono (sInf_le (i : T).2) + +theorem regularSpace_iInf {ι X} {t : ι → TopologicalSpace X} (h : ∀ i, @RegularSpace X (t i)) : + @RegularSpace X (iInf t) := + regularSpace_sInf <| forall_mem_range.mpr h + +theorem RegularSpace.inf {X} {t₁ t₂ : TopologicalSpace X} (h₁ : @RegularSpace X t₁) + (h₂ : @RegularSpace X t₂) : @RegularSpace X (t₁ ⊓ t₂) := by + rw [inf_eq_iInf] + exact regularSpace_iInf (Bool.forall_bool.2 ⟨h₂, h₁⟩) + +instance {p : X → Prop} : RegularSpace (Subtype p) := + IsEmbedding.subtypeVal.isInducing.regularSpace + +instance [TopologicalSpace Y] [RegularSpace Y] : RegularSpace (X × Y) := + (regularSpace_induced (@Prod.fst X Y)).inf (regularSpace_induced (@Prod.snd X Y)) + +instance {ι : Type*} {X : ι → Type*} [∀ i, TopologicalSpace (X i)] [∀ i, RegularSpace (X i)] : + RegularSpace (∀ i, X i) := + regularSpace_iInf fun _ => regularSpace_induced _ + +/-- In a regular space, if a compact set and a closed set are disjoint, then they have disjoint +neighborhoods. -/ +lemma SeparatedNhds.of_isCompact_isClosed {s t : Set X} + (hs : IsCompact s) (ht : IsClosed t) (hst : Disjoint s t) : SeparatedNhds s t := by + simpa only [separatedNhds_iff_disjoint, hs.disjoint_nhdsSet_left, disjoint_nhds_nhdsSet, + ht.closure_eq, disjoint_left] using hst + +end + +/-- This technique to witness `HasSeparatingCover` in regular Lindelöf topological spaces +will be used to prove regular Lindelöf spaces are normal. -/ +lemma IsClosed.HasSeparatingCover {s t : Set X} [LindelofSpace X] [RegularSpace X] + (s_cl : IsClosed s) (t_cl : IsClosed t) (st_dis : Disjoint s t) : HasSeparatingCover s t := by + -- `IsLindelof.indexed_countable_subcover` requires the space be Nonempty + rcases isEmpty_or_nonempty X with empty_X | nonempty_X + · rw [subset_eq_empty (t := s) (fun ⦃_⦄ _ ↦ trivial) (univ_eq_empty_iff.mpr empty_X)] + exact hasSeparatingCovers_iff_separatedNhds.mpr (SeparatedNhds.empty_left t) |>.1 + -- This is almost `HasSeparatingCover`, but is not countable. We define for all `a : X` for use + -- with `IsLindelof.indexed_countable_subcover` momentarily. + have (a : X) : ∃ n : Set X, IsOpen n ∧ Disjoint (closure n) t ∧ (a ∈ s → a ∈ n) := by + wlog ains : a ∈ s + · exact ⟨∅, isOpen_empty, SeparatedNhds.empty_left t |>.disjoint_closure_left, fun a ↦ ains a⟩ + obtain ⟨n, nna, ncl, nsubkc⟩ := ((regularSpace_TFAE X).out 0 3 :).mp ‹RegularSpace X› a tᶜ <| + t_cl.compl_mem_nhds (disjoint_left.mp st_dis ains) + exact + ⟨interior n, + isOpen_interior, + disjoint_left.mpr fun ⦃_⦄ ain ↦ + nsubkc <| (IsClosed.closure_subset_iff ncl).mpr interior_subset ain, + fun _ ↦ mem_interior_iff_mem_nhds.mpr nna⟩ + -- By Lindelöf, we may obtain a countable subcover witnessing `HasSeparatingCover` + choose u u_open u_dis u_nhd using this + obtain ⟨f, f_cov⟩ := s_cl.isLindelof.indexed_countable_subcover + u u_open (fun a ainh ↦ mem_iUnion.mpr ⟨a, u_nhd a ainh⟩) + exact ⟨u ∘ f, f_cov, fun n ↦ ⟨u_open (f n), u_dis (f n)⟩⟩ + + +end RegularSpace + +section LocallyCompactRegularSpace + +/-- In a (possibly non-Hausdorff) locally compact regular space, for every containment `K ⊆ U` of + a compact set `K` in an open set `U`, there is a compact closed neighborhood `L` + such that `K ⊆ L ⊆ U`: equivalently, there is a compact closed set `L` such + that `K ⊆ interior L` and `L ⊆ U`. -/ +theorem exists_compact_closed_between [LocallyCompactSpace X] [RegularSpace X] + {K U : Set X} (hK : IsCompact K) (hU : IsOpen U) (h_KU : K ⊆ U) : + ∃ L, IsCompact L ∧ IsClosed L ∧ K ⊆ interior L ∧ L ⊆ U := + let ⟨L, L_comp, KL, LU⟩ := exists_compact_between hK hU h_KU + ⟨closure L, L_comp.closure, isClosed_closure, KL.trans <| interior_mono subset_closure, + L_comp.closure_subset_of_isOpen hU LU⟩ + +/-- In a locally compact regular space, given a compact set `K` inside an open set `U`, we can find +an open set `V` between these sets with compact closure: `K ⊆ V` and the closure of `V` is +inside `U`. -/ +theorem exists_open_between_and_isCompact_closure [LocallyCompactSpace X] [RegularSpace X] + {K U : Set X} (hK : IsCompact K) (hU : IsOpen U) (hKU : K ⊆ U) : + ∃ V, IsOpen V ∧ K ⊆ V ∧ closure V ⊆ U ∧ IsCompact (closure V) := by + rcases exists_compact_closed_between hK hU hKU with ⟨L, L_compact, L_closed, KL, LU⟩ + have A : closure (interior L) ⊆ L := by + apply (closure_mono interior_subset).trans (le_of_eq L_closed.closure_eq) + refine ⟨interior L, isOpen_interior, KL, A.trans LU, ?_⟩ + exact L_compact.closure_of_subset interior_subset + +end LocallyCompactRegularSpace + +section T25 + +/-- A T₂.₅ space, also known as a Urysohn space, is a topological space + where for every pair `x ≠ y`, there are two open sets, with the intersection of closures + empty, one containing `x` and the other `y` . -/ +class T25Space (X : Type u) [TopologicalSpace X] : Prop where + /-- Given two distinct points in a T₂.₅ space, their filters of closed neighborhoods are + disjoint. -/ + t2_5 : ∀ ⦃x y : X⦄, x ≠ y → Disjoint ((𝓝 x).lift' closure) ((𝓝 y).lift' closure) + +@[simp] +theorem disjoint_lift'_closure_nhds [T25Space X] {x y : X} : + Disjoint ((𝓝 x).lift' closure) ((𝓝 y).lift' closure) ↔ x ≠ y := + ⟨fun h hxy => by simp [hxy, nhds_neBot.ne] at h, fun h => T25Space.t2_5 h⟩ + +-- see Note [lower instance priority] +instance (priority := 100) T25Space.t2Space [T25Space X] : T2Space X := + t2Space_iff_disjoint_nhds.2 fun _ _ hne => + (disjoint_lift'_closure_nhds.2 hne).mono (le_lift'_closure _) (le_lift'_closure _) + +theorem exists_nhds_disjoint_closure [T25Space X] {x y : X} (h : x ≠ y) : + ∃ s ∈ 𝓝 x, ∃ t ∈ 𝓝 y, Disjoint (closure s) (closure t) := + ((𝓝 x).basis_sets.lift'_closure.disjoint_iff (𝓝 y).basis_sets.lift'_closure).1 <| + disjoint_lift'_closure_nhds.2 h + +theorem exists_open_nhds_disjoint_closure [T25Space X] {x y : X} (h : x ≠ y) : + ∃ u : Set X, + x ∈ u ∧ IsOpen u ∧ ∃ v : Set X, y ∈ v ∧ IsOpen v ∧ Disjoint (closure u) (closure v) := by + simpa only [exists_prop, and_assoc] using + ((nhds_basis_opens x).lift'_closure.disjoint_iff (nhds_basis_opens y).lift'_closure).1 + (disjoint_lift'_closure_nhds.2 h) + +theorem T25Space.of_injective_continuous [TopologicalSpace Y] [T25Space Y] {f : X → Y} + (hinj : Injective f) (hcont : Continuous f) : T25Space X where + t2_5 x y hne := (tendsto_lift'_closure_nhds hcont x).disjoint (t2_5 <| hinj.ne hne) + (tendsto_lift'_closure_nhds hcont y) + +theorem Topology.IsEmbedding.t25Space [TopologicalSpace Y] [T25Space Y] {f : X → Y} + (hf : IsEmbedding f) : T25Space X := + .of_injective_continuous hf.injective hf.continuous + +@[deprecated (since := "2024-10-26")] +alias Embedding.t25Space := IsEmbedding.t25Space + +instance Subtype.instT25Space [T25Space X] {p : X → Prop} : T25Space {x // p x} := + IsEmbedding.subtypeVal.t25Space + +end T25 + +section T3 + +/-- A T₃ space is a T₀ space which is a regular space. Any T₃ space is a T₁ space, a T₂ space, and +a T₂.₅ space. -/ +class T3Space (X : Type u) [TopologicalSpace X] extends T0Space X, RegularSpace X : Prop + +instance (priority := 90) instT3Space [T0Space X] [RegularSpace X] : T3Space X := ⟨⟩ + +theorem RegularSpace.t3Space_iff_t0Space [RegularSpace X] : T3Space X ↔ T0Space X := by + constructor <;> intro <;> infer_instance + +-- see Note [lower instance priority] +instance (priority := 100) T3Space.t25Space [T3Space X] : T25Space X := by + refine ⟨fun x y hne => ?_⟩ + rw [lift'_nhds_closure, lift'_nhds_closure] + have : x ∉ closure {y} ∨ y ∉ closure {x} := + (t0Space_iff_or_not_mem_closure X).mp inferInstance hne + simp only [← disjoint_nhds_nhdsSet, nhdsSet_singleton] at this + exact this.elim id fun h => h.symm + +protected theorem Topology.IsEmbedding.t3Space [TopologicalSpace Y] [T3Space Y] {f : X → Y} + (hf : IsEmbedding f) : T3Space X := + { toT0Space := hf.t0Space + toRegularSpace := hf.isInducing.regularSpace } + +@[deprecated (since := "2024-10-26")] +alias Embedding.t3Space := IsEmbedding.t3Space + +instance Subtype.t3Space [T3Space X] {p : X → Prop} : T3Space (Subtype p) := + IsEmbedding.subtypeVal.t3Space + +instance ULift.instT3Space [T3Space X] : T3Space (ULift X) := + IsEmbedding.uliftDown.t3Space + +instance [TopologicalSpace Y] [T3Space X] [T3Space Y] : T3Space (X × Y) := ⟨⟩ + +instance {ι : Type*} {X : ι → Type*} [∀ i, TopologicalSpace (X i)] [∀ i, T3Space (X i)] : + T3Space (∀ i, X i) := ⟨⟩ + +/-- Given two points `x ≠ y`, we can find neighbourhoods `x ∈ V₁ ⊆ U₁` and `y ∈ V₂ ⊆ U₂`, +with the `Vₖ` closed and the `Uₖ` open, such that the `Uₖ` are disjoint. -/ +theorem disjoint_nested_nhds [T3Space X] {x y : X} (h : x ≠ y) : + ∃ U₁ ∈ 𝓝 x, ∃ V₁ ∈ 𝓝 x, ∃ U₂ ∈ 𝓝 y, ∃ V₂ ∈ 𝓝 y, + IsClosed V₁ ∧ IsClosed V₂ ∧ IsOpen U₁ ∧ IsOpen U₂ ∧ V₁ ⊆ U₁ ∧ V₂ ⊆ U₂ ∧ Disjoint U₁ U₂ := by + rcases t2_separation h with ⟨U₁, U₂, U₁_op, U₂_op, x_in, y_in, H⟩ + rcases exists_mem_nhds_isClosed_subset (U₁_op.mem_nhds x_in) with ⟨V₁, V₁_in, V₁_closed, h₁⟩ + rcases exists_mem_nhds_isClosed_subset (U₂_op.mem_nhds y_in) with ⟨V₂, V₂_in, V₂_closed, h₂⟩ + exact ⟨U₁, mem_of_superset V₁_in h₁, V₁, V₁_in, U₂, mem_of_superset V₂_in h₂, V₂, V₂_in, + V₁_closed, V₂_closed, U₁_op, U₂_op, h₁, h₂, H⟩ + +open SeparationQuotient + +/-- The `SeparationQuotient` of a regular space is a T₃ space. -/ +instance [RegularSpace X] : T3Space (SeparationQuotient X) where + regular {s a} hs ha := by + rcases surjective_mk a with ⟨a, rfl⟩ + rw [← disjoint_comap_iff surjective_mk, comap_mk_nhds_mk, comap_mk_nhdsSet] + exact RegularSpace.regular (hs.preimage continuous_mk) ha + +end T3 + +section NormalSpace + +/-- A topological space is said to be a *normal space* if any two disjoint closed sets +have disjoint open neighborhoods. -/ +class NormalSpace (X : Type u) [TopologicalSpace X] : Prop where + /-- Two disjoint sets in a normal space admit disjoint neighbourhoods. -/ + normal : ∀ s t : Set X, IsClosed s → IsClosed t → Disjoint s t → SeparatedNhds s t + +theorem normal_separation [NormalSpace X] {s t : Set X} (H1 : IsClosed s) (H2 : IsClosed t) + (H3 : Disjoint s t) : SeparatedNhds s t := + NormalSpace.normal s t H1 H2 H3 + +theorem disjoint_nhdsSet_nhdsSet [NormalSpace X] {s t : Set X} (hs : IsClosed s) (ht : IsClosed t) + (hd : Disjoint s t) : Disjoint (𝓝ˢ s) (𝓝ˢ t) := + (normal_separation hs ht hd).disjoint_nhdsSet + +theorem normal_exists_closure_subset [NormalSpace X] {s t : Set X} (hs : IsClosed s) (ht : IsOpen t) + (hst : s ⊆ t) : ∃ u, IsOpen u ∧ s ⊆ u ∧ closure u ⊆ t := by + have : Disjoint s tᶜ := Set.disjoint_left.mpr fun x hxs hxt => hxt (hst hxs) + rcases normal_separation hs (isClosed_compl_iff.2 ht) this with + ⟨s', t', hs', ht', hss', htt', hs't'⟩ + refine ⟨s', hs', hss', Subset.trans (closure_minimal ?_ (isClosed_compl_iff.2 ht')) + (compl_subset_comm.1 htt')⟩ + exact fun x hxs hxt => hs't'.le_bot ⟨hxs, hxt⟩ + +/-- If the codomain of a closed embedding is a normal space, then so is the domain. -/ +protected theorem Topology.IsClosedEmbedding.normalSpace [TopologicalSpace Y] [NormalSpace Y] + {f : X → Y} (hf : IsClosedEmbedding f) : NormalSpace X where + normal s t hs ht hst := by + have H : SeparatedNhds (f '' s) (f '' t) := + NormalSpace.normal (f '' s) (f '' t) (hf.isClosedMap s hs) (hf.isClosedMap t ht) + (disjoint_image_of_injective hf.injective hst) + exact (H.preimage hf.continuous).mono (subset_preimage_image _ _) (subset_preimage_image _ _) + +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.normalSpace := IsClosedEmbedding.normalSpace + +instance (priority := 100) NormalSpace.of_compactSpace_r1Space [CompactSpace X] [R1Space X] : + NormalSpace X where + normal _s _t hs ht := .of_isCompact_isCompact_isClosed hs.isCompact ht.isCompact ht + +set_option pp.universes true in +/-- A regular topological space with a Lindelöf topology is a normal space. A consequence of e.g. +Corollaries 20.8 and 20.10 of [Willard's *General Topology*][zbMATH02107988] (without the +assumption of Hausdorff). -/ +instance (priority := 100) NormalSpace.of_regularSpace_lindelofSpace + [RegularSpace X] [LindelofSpace X] : NormalSpace X where + normal _ _ hcl kcl hkdis := + hasSeparatingCovers_iff_separatedNhds.mp + ⟨hcl.HasSeparatingCover kcl hkdis, kcl.HasSeparatingCover hcl (Disjoint.symm hkdis)⟩ + +instance (priority := 100) NormalSpace.of_regularSpace_secondCountableTopology + [RegularSpace X] [SecondCountableTopology X] : NormalSpace X := + of_regularSpace_lindelofSpace + +end NormalSpace + +section Normality + +/-- A T₄ space is a normal T₁ space. -/ +class T4Space (X : Type u) [TopologicalSpace X] extends T1Space X, NormalSpace X : Prop + +instance (priority := 100) [T1Space X] [NormalSpace X] : T4Space X := ⟨⟩ + +-- see Note [lower instance priority] +instance (priority := 100) T4Space.t3Space [T4Space X] : T3Space X where + regular hs hxs := by simpa only [nhdsSet_singleton] using (normal_separation hs isClosed_singleton + (disjoint_singleton_right.mpr hxs)).disjoint_nhdsSet + +/-- If the codomain of a closed embedding is a T₄ space, then so is the domain. -/ +protected theorem Topology.IsClosedEmbedding.t4Space [TopologicalSpace Y] [T4Space Y] {f : X → Y} + (hf : IsClosedEmbedding f) : T4Space X where + toT1Space := hf.isEmbedding.t1Space + toNormalSpace := hf.normalSpace + +@[deprecated (since := "2024-10-20")] +alias ClosedEmbedding.t4Space := IsClosedEmbedding.t4Space + +instance ULift.instT4Space [T4Space X] : T4Space (ULift X) := IsClosedEmbedding.uliftDown.t4Space + +namespace SeparationQuotient + +/-- The `SeparationQuotient` of a normal space is a normal space. -/ +instance [NormalSpace X] : NormalSpace (SeparationQuotient X) where + normal s t hs ht hd := separatedNhds_iff_disjoint.2 <| by + rw [← disjoint_comap_iff surjective_mk, comap_mk_nhdsSet, comap_mk_nhdsSet] + exact disjoint_nhdsSet_nhdsSet (hs.preimage continuous_mk) (ht.preimage continuous_mk) + (hd.preimage mk) + +end SeparationQuotient + +variable (X) + +end Normality + +section CompletelyNormal + +/-- A topological space `X` is a *completely normal space* provided that for any two sets `s`, `t` +such that if both `closure s` is disjoint with `t`, and `s` is disjoint with `closure t`, +then there exist disjoint neighbourhoods of `s` and `t`. -/ +class CompletelyNormalSpace (X : Type u) [TopologicalSpace X] : Prop where + /-- If `closure s` is disjoint with `t`, and `s` is disjoint with `closure t`, then `s` and `t` + admit disjoint neighbourhoods. -/ + completely_normal : + ∀ ⦃s t : Set X⦄, Disjoint (closure s) t → Disjoint s (closure t) → Disjoint (𝓝ˢ s) (𝓝ˢ t) + +export CompletelyNormalSpace (completely_normal) + +-- see Note [lower instance priority] +/-- A completely normal space is a normal space. -/ +instance (priority := 100) CompletelyNormalSpace.toNormalSpace + [CompletelyNormalSpace X] : NormalSpace X where + normal s t hs ht hd := separatedNhds_iff_disjoint.2 <| + completely_normal (by rwa [hs.closure_eq]) (by rwa [ht.closure_eq]) + +theorem Topology.IsEmbedding.completelyNormalSpace [TopologicalSpace Y] [CompletelyNormalSpace Y] + {e : X → Y} (he : IsEmbedding e) : CompletelyNormalSpace X := by + refine ⟨fun s t hd₁ hd₂ => ?_⟩ + simp only [he.isInducing.nhdsSet_eq_comap] + refine disjoint_comap (completely_normal ?_ ?_) + · rwa [← subset_compl_iff_disjoint_left, image_subset_iff, preimage_compl, + ← he.closure_eq_preimage_closure_image, subset_compl_iff_disjoint_left] + · rwa [← subset_compl_iff_disjoint_right, image_subset_iff, preimage_compl, + ← he.closure_eq_preimage_closure_image, subset_compl_iff_disjoint_right] + +@[deprecated (since := "2024-10-26")] +alias Embedding.completelyNormalSpace := IsEmbedding.completelyNormalSpace + +/-- A subspace of a completely normal space is a completely normal space. -/ +instance [CompletelyNormalSpace X] {p : X → Prop} : CompletelyNormalSpace { x // p x } := + IsEmbedding.subtypeVal.completelyNormalSpace + +instance ULift.instCompletelyNormalSpace [CompletelyNormalSpace X] : + CompletelyNormalSpace (ULift X) := + IsEmbedding.uliftDown.completelyNormalSpace + +/-- A T₅ space is a completely normal T₁ space. -/ +class T5Space (X : Type u) [TopologicalSpace X] extends T1Space X, CompletelyNormalSpace X : Prop + +theorem Topology.IsEmbedding.t5Space [TopologicalSpace Y] [T5Space Y] {e : X → Y} + (he : IsEmbedding e) : T5Space X where + __ := he.t1Space + completely_normal := by + have := he.completelyNormalSpace + exact completely_normal + +@[deprecated (since := "2024-10-26")] +alias Embedding.t5Space := IsEmbedding.t5Space + +-- see Note [lower instance priority] +/-- A `T₅` space is a `T₄` space. -/ +instance (priority := 100) T5Space.toT4Space [T5Space X] : T4Space X where + -- follows from type-class inference + +/-- A subspace of a T₅ space is a T₅ space. -/ +instance [T5Space X] {p : X → Prop} : T5Space { x // p x } := + IsEmbedding.subtypeVal.t5Space + +instance ULift.instT5Space [T5Space X] : T5Space (ULift X) := + IsEmbedding.uliftDown.t5Space + +open SeparationQuotient + +/-- The `SeparationQuotient` of a completely normal R₀ space is a T₅ space. -/ +instance [CompletelyNormalSpace X] [R0Space X] : T5Space (SeparationQuotient X) where + t1 := by + rwa [((t1Space_TFAE (SeparationQuotient X)).out 1 0 :), SeparationQuotient.t1Space_iff] + completely_normal s t hd₁ hd₂ := by + rw [← disjoint_comap_iff surjective_mk, comap_mk_nhdsSet, comap_mk_nhdsSet] + apply completely_normal <;> rw [← preimage_mk_closure] + exacts [hd₁.preimage mk, hd₂.preimage mk] + +end CompletelyNormal + +/-- In a compact T₂ space, the connected component of a point equals the intersection of all +its clopen neighbourhoods. -/ +theorem connectedComponent_eq_iInter_isClopen [T2Space X] [CompactSpace X] (x : X) : + connectedComponent x = ⋂ s : { s : Set X // IsClopen s ∧ x ∈ s }, s := by + apply Subset.antisymm connectedComponent_subset_iInter_isClopen + -- Reduce to showing that the clopen intersection is connected. + refine IsPreconnected.subset_connectedComponent ?_ (mem_iInter.2 fun s => s.2.2) + -- We do this by showing that any disjoint cover by two closed sets implies + -- that one of these closed sets must contain our whole thing. + -- To reduce to the case where the cover is disjoint on all of `X` we need that `s` is closed + have hs : @IsClosed X _ (⋂ s : { s : Set X // IsClopen s ∧ x ∈ s }, s) := + isClosed_iInter fun s => s.2.1.1 + rw [isPreconnected_iff_subset_of_fully_disjoint_closed hs] + intro a b ha hb hab ab_disj + -- Since our space is normal, we get two larger disjoint open sets containing the disjoint + -- closed sets. If we can show that our intersection is a subset of any of these we can then + -- "descend" this to show that it is a subset of either a or b. + rcases normal_separation ha hb ab_disj with ⟨u, v, hu, hv, hau, hbv, huv⟩ + obtain ⟨s, H⟩ : ∃ s : Set X, IsClopen s ∧ x ∈ s ∧ s ⊆ u ∪ v := by + /- Now we find a clopen set `s` around `x`, contained in `u ∪ v`. We utilize the fact that + `X \ u ∪ v` will be compact, so there must be some finite intersection of clopen neighbourhoods + of `X` disjoint to it, but a finite intersection of clopen sets is clopen, + so we let this be our `s`. -/ + have H1 := (hu.union hv).isClosed_compl.isCompact.inter_iInter_nonempty + (fun s : { s : Set X // IsClopen s ∧ x ∈ s } => s) fun s => s.2.1.1 + rw [← not_disjoint_iff_nonempty_inter, imp_not_comm, not_forall] at H1 + cases' H1 (disjoint_compl_left_iff_subset.2 <| hab.trans <| union_subset_union hau hbv) + with si H2 + refine ⟨⋂ U ∈ si, Subtype.val U, ?_, ?_, ?_⟩ + · exact isClopen_biInter_finset fun s _ => s.2.1 + · exact mem_iInter₂.2 fun s _ => s.2.2 + · rwa [← disjoint_compl_left_iff_subset, disjoint_iff_inter_eq_empty, + ← not_nonempty_iff_eq_empty] + -- So, we get a disjoint decomposition `s = s ∩ u ∪ s ∩ v` of clopen sets. The intersection of all + -- clopen neighbourhoods will then lie in whichever of u or v x lies in and hence will be a subset + -- of either a or b. + · have H1 := isClopen_inter_of_disjoint_cover_clopen H.1 H.2.2 hu hv huv + rw [union_comm] at H + have H2 := isClopen_inter_of_disjoint_cover_clopen H.1 H.2.2 hv hu huv.symm + by_cases hxu : x ∈ u <;> [left; right] + -- The x ∈ u case. + · suffices ⋂ s : { s : Set X // IsClopen s ∧ x ∈ s }, ↑s ⊆ u + from Disjoint.left_le_of_le_sup_right hab (huv.mono this hbv) + · apply Subset.trans _ s.inter_subset_right + exact iInter_subset (fun s : { s : Set X // IsClopen s ∧ x ∈ s } => s.1) + ⟨s ∩ u, H1, mem_inter H.2.1 hxu⟩ + -- If x ∉ u, we get x ∈ v since x ∈ u ∪ v. The rest is then like the x ∈ u case. + · have h1 : x ∈ v := + (hab.trans (union_subset_union hau hbv) (mem_iInter.2 fun i => i.2.2)).resolve_left hxu + suffices ⋂ s : { s : Set X // IsClopen s ∧ x ∈ s }, ↑s ⊆ v + from (huv.symm.mono this hau).left_le_of_le_sup_left hab + · refine Subset.trans ?_ s.inter_subset_right + exact iInter_subset (fun s : { s : Set X // IsClopen s ∧ x ∈ s } => s.1) + ⟨s ∩ v, H2, mem_inter H.2.1 h1⟩ + +/-- `ConnectedComponents X` is Hausdorff when `X` is Hausdorff and compact -/ +instance ConnectedComponents.t2 [T2Space X] [CompactSpace X] : T2Space (ConnectedComponents X) := by + -- Proof follows that of: https://stacks.math.columbia.edu/tag/0900 + -- Fix 2 distinct connected components, with points a and b + refine ⟨ConnectedComponents.surjective_coe.forall₂.2 fun a b ne => ?_⟩ + rw [ConnectedComponents.coe_ne_coe] at ne + have h := connectedComponent_disjoint ne + -- write ↑b as the intersection of all clopen subsets containing it + rw [connectedComponent_eq_iInter_isClopen b, disjoint_iff_inter_eq_empty] at h + -- Now we show that this can be reduced to some clopen containing `↑b` being disjoint to `↑a` + obtain ⟨U, V, hU, ha, hb, rfl⟩ : ∃ (U : Set X) (V : Set (ConnectedComponents X)), + IsClopen U ∧ connectedComponent a ∩ U = ∅ ∧ connectedComponent b ⊆ U ∧ (↑) ⁻¹' V = U := by + have h := + (isClosed_connectedComponent (α := X)).isCompact.elim_finite_subfamily_closed + _ (fun s : { s : Set X // IsClopen s ∧ b ∈ s } => s.2.1.1) h + cases' h with fin_a ha + -- This clopen and its complement will separate the connected components of `a` and `b` + set U : Set X := ⋂ (i : { s // IsClopen s ∧ b ∈ s }) (_ : i ∈ fin_a), i + have hU : IsClopen U := isClopen_biInter_finset fun i _ => i.2.1 + exact ⟨U, (↑) '' U, hU, ha, subset_iInter₂ fun s _ => s.2.1.connectedComponent_subset s.2.2, + (connectedComponents_preimage_image U).symm ▸ hU.biUnion_connectedComponent_eq⟩ + rw [ConnectedComponents.isQuotientMap_coe.isClopen_preimage] at hU + refine ⟨Vᶜ, V, hU.compl.isOpen, hU.isOpen, ?_, hb mem_connectedComponent, disjoint_compl_left⟩ + exact fun h => flip Set.Nonempty.ne_empty ha ⟨a, mem_connectedComponent, h⟩ diff --git a/Mathlib/Topology/Separation/SeparatedNhds.lean b/Mathlib/Topology/Separation/SeparatedNhds.lean new file mode 100644 index 0000000000000..0466117eb2ddb --- /dev/null +++ b/Mathlib/Topology/Separation/SeparatedNhds.lean @@ -0,0 +1,160 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Mario Carneiro +-/ +import Mathlib.Topology.NhdsSet + +/-! +# Separated neighbourhoods + +This file defines the predicates `SeparatedNhds` and `HasSeparatingCover`, which are used in +formulating separation axioms for topological spaces. + +## Main definitions + +* `SeparatedNhds`: Two `Set`s are separated by neighbourhoods if they are contained in disjoint + open sets. +* `HasSeparatingCover`: A set has a countable cover that can be used with + `hasSeparatingCovers_iff_separatedNhds` to witness when two `Set`s have `SeparatedNhds`. + +## References + +* +* [Willard's *General Topology*][zbMATH02107988] +-/ + +open Function Set Filter Topology TopologicalSpace + +universe u v + +variable {X : Type*} {Y : Type*} [TopologicalSpace X] + +section Separation + +/-- +`SeparatedNhds` is a predicate on pairs of sub`Set`s of a topological space. It holds if the two +sub`Set`s are contained in disjoint open sets. +-/ +def SeparatedNhds : Set X → Set X → Prop := fun s t : Set X => + ∃ U V : Set X, IsOpen U ∧ IsOpen V ∧ s ⊆ U ∧ t ⊆ V ∧ Disjoint U V + +theorem separatedNhds_iff_disjoint {s t : Set X} : SeparatedNhds s t ↔ Disjoint (𝓝ˢ s) (𝓝ˢ t) := by + simp only [(hasBasis_nhdsSet s).disjoint_iff (hasBasis_nhdsSet t), SeparatedNhds, exists_prop, ← + exists_and_left, and_assoc, and_comm, and_left_comm] + +alias ⟨SeparatedNhds.disjoint_nhdsSet, _⟩ := separatedNhds_iff_disjoint + +/-- `HasSeparatingCover`s can be useful witnesses for `SeparatedNhds`. -/ +def HasSeparatingCover : Set X → Set X → Prop := fun s t ↦ + ∃ u : ℕ → Set X, s ⊆ ⋃ n, u n ∧ ∀ n, IsOpen (u n) ∧ Disjoint (closure (u n)) t + +/-- Used to prove that a regular topological space with Lindelöf topology is a normal space, +and a perfectly normal space is a completely normal space. -/ +theorem hasSeparatingCovers_iff_separatedNhds {s t : Set X} : + HasSeparatingCover s t ∧ HasSeparatingCover t s ↔ SeparatedNhds s t := by + constructor + · rintro ⟨⟨u, u_cov, u_props⟩, ⟨v, v_cov, v_props⟩⟩ + have open_lemma : ∀ (u₀ a : ℕ → Set X), (∀ n, IsOpen (u₀ n)) → + IsOpen (⋃ n, u₀ n \ closure (a n)) := fun _ _ u₀i_open ↦ + isOpen_iUnion fun i ↦ (u₀i_open i).sdiff isClosed_closure + have cover_lemma : ∀ (h₀ : Set X) (u₀ v₀ : ℕ → Set X), + (h₀ ⊆ ⋃ n, u₀ n) → (∀ n, Disjoint (closure (v₀ n)) h₀) → + (h₀ ⊆ ⋃ n, u₀ n \ closure (⋃ m ≤ n, v₀ m)) := + fun h₀ u₀ v₀ h₀_cov dis x xinh ↦ by + rcases h₀_cov xinh with ⟨un , ⟨n, rfl⟩ , xinun⟩ + simp only [mem_iUnion] + refine ⟨n, xinun, ?_⟩ + simp_all only [closure_iUnion₂_le_nat, disjoint_right, mem_setOf_eq, mem_iUnion, + exists_false, exists_const, not_false_eq_true] + refine + ⟨⋃ n : ℕ, u n \ (closure (⋃ m ≤ n, v m)), + ⋃ n : ℕ, v n \ (closure (⋃ m ≤ n, u m)), + open_lemma u (fun n ↦ ⋃ m ≤ n, v m) (fun n ↦ (u_props n).1), + open_lemma v (fun n ↦ ⋃ m ≤ n, u m) (fun n ↦ (v_props n).1), + cover_lemma s u v u_cov (fun n ↦ (v_props n).2), + cover_lemma t v u v_cov (fun n ↦ (u_props n).2), + ?_⟩ + rw [Set.disjoint_left] + rintro x ⟨un, ⟨n, rfl⟩, xinun⟩ + suffices ∀ (m : ℕ), x ∈ v m → x ∈ closure (⋃ m' ∈ {m' | m' ≤ m}, u m') by simpa + intro m xinvm + have n_le_m : n ≤ m := by + by_contra m_gt_n + exact xinun.2 (subset_closure (mem_biUnion (le_of_lt (not_le.mp m_gt_n)) xinvm)) + exact subset_closure (mem_biUnion n_le_m xinun.1) + · rintro ⟨U, V, U_open, V_open, h_sub_U, k_sub_V, UV_dis⟩ + exact + ⟨⟨fun _ ↦ U, + h_sub_U.trans (iUnion_const U).symm.subset, + fun _ ↦ + ⟨U_open, disjoint_of_subset (fun ⦃a⦄ a ↦ a) k_sub_V (UV_dis.closure_left V_open)⟩⟩, + ⟨fun _ ↦ V, + k_sub_V.trans (iUnion_const V).symm.subset, + fun _ ↦ + ⟨V_open, disjoint_of_subset (fun ⦃a⦄ a ↦ a) h_sub_U (UV_dis.closure_right U_open).symm⟩⟩⟩ + +theorem Set.hasSeparatingCover_empty_left (s : Set X) : HasSeparatingCover ∅ s := + ⟨fun _ ↦ ∅, empty_subset (⋃ _, ∅), + fun _ ↦ ⟨isOpen_empty, by simp only [closure_empty, empty_disjoint]⟩⟩ + +theorem Set.hasSeparatingCover_empty_right (s : Set X) : HasSeparatingCover s ∅ := + ⟨fun _ ↦ univ, (subset_univ s).trans univ.iUnion_const.symm.subset, + fun _ ↦ ⟨isOpen_univ, by apply disjoint_empty⟩⟩ + +theorem HasSeparatingCover.mono {s₁ s₂ t₁ t₂ : Set X} (sc_st : HasSeparatingCover s₂ t₂) + (s_sub : s₁ ⊆ s₂) (t_sub : t₁ ⊆ t₂) : HasSeparatingCover s₁ t₁ := by + obtain ⟨u, u_cov, u_props⟩ := sc_st + exact + ⟨u, + s_sub.trans u_cov, + fun n ↦ + ⟨(u_props n).1, + disjoint_of_subset (fun ⦃_⦄ a ↦ a) t_sub (u_props n).2⟩⟩ + +namespace SeparatedNhds + +variable {s s₁ s₂ t t₁ t₂ u : Set X} + +@[symm] +theorem symm : SeparatedNhds s t → SeparatedNhds t s := fun ⟨U, V, oU, oV, aU, bV, UV⟩ => + ⟨V, U, oV, oU, bV, aU, Disjoint.symm UV⟩ + +theorem comm (s t : Set X) : SeparatedNhds s t ↔ SeparatedNhds t s := + ⟨symm, symm⟩ + +theorem preimage [TopologicalSpace Y] {f : X → Y} {s t : Set Y} (h : SeparatedNhds s t) + (hf : Continuous f) : SeparatedNhds (f ⁻¹' s) (f ⁻¹' t) := + let ⟨U, V, oU, oV, sU, tV, UV⟩ := h + ⟨f ⁻¹' U, f ⁻¹' V, oU.preimage hf, oV.preimage hf, preimage_mono sU, preimage_mono tV, + UV.preimage f⟩ + +protected theorem disjoint (h : SeparatedNhds s t) : Disjoint s t := + let ⟨_, _, _, _, hsU, htV, hd⟩ := h; hd.mono hsU htV + +theorem disjoint_closure_left (h : SeparatedNhds s t) : Disjoint (closure s) t := + let ⟨_U, _V, _, hV, hsU, htV, hd⟩ := h + (hd.closure_left hV).mono (closure_mono hsU) htV + +theorem disjoint_closure_right (h : SeparatedNhds s t) : Disjoint s (closure t) := + h.symm.disjoint_closure_left.symm + +@[simp] theorem empty_right (s : Set X) : SeparatedNhds s ∅ := + ⟨_, _, isOpen_univ, isOpen_empty, fun a _ => mem_univ a, Subset.rfl, disjoint_empty _⟩ + +@[simp] theorem empty_left (s : Set X) : SeparatedNhds ∅ s := + (empty_right _).symm + +theorem mono (h : SeparatedNhds s₂ t₂) (hs : s₁ ⊆ s₂) (ht : t₁ ⊆ t₂) : SeparatedNhds s₁ t₁ := + let ⟨U, V, hU, hV, hsU, htV, hd⟩ := h + ⟨U, V, hU, hV, hs.trans hsU, ht.trans htV, hd⟩ + +theorem union_left : SeparatedNhds s u → SeparatedNhds t u → SeparatedNhds (s ∪ t) u := by + simpa only [separatedNhds_iff_disjoint, nhdsSet_union, disjoint_sup_left] using And.intro + +theorem union_right (ht : SeparatedNhds s t) (hu : SeparatedNhds s u) : SeparatedNhds s (t ∪ u) := + (ht.symm.union_left hu.symm).symm + +end SeparatedNhds + +end Separation diff --git a/Mathlib/Topology/ShrinkingLemma.lean b/Mathlib/Topology/ShrinkingLemma.lean index 65f6b72889730..2c85b7383164c 100644 --- a/Mathlib/Topology/ShrinkingLemma.lean +++ b/Mathlib/Topology/ShrinkingLemma.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Reid Barton -/ -import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.Separation.Regular /-! # The shrinking lemma diff --git a/Mathlib/Topology/UniformSpace/OfCompactT2.lean b/Mathlib/Topology/UniformSpace/OfCompactT2.lean index c9d49e9474570..4c3a279c43064 100644 --- a/Mathlib/Topology/UniformSpace/OfCompactT2.lean +++ b/Mathlib/Topology/UniformSpace/OfCompactT2.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Patrick Massot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Yury Kudryashov -/ -import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.Separation.Regular import Mathlib.Topology.UniformSpace.Basic /-! diff --git a/Mathlib/Topology/UniformSpace/Separation.lean b/Mathlib/Topology/UniformSpace/Separation.lean index 9b3a324dc67df..4fc507fc540cb 100644 --- a/Mathlib/Topology/UniformSpace/Separation.lean +++ b/Mathlib/Topology/UniformSpace/Separation.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Patrick Massot, Yury Kudryashov -/ import Mathlib.Tactic.ApplyFun -import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.Separation.Regular import Mathlib.Topology.UniformSpace.Basic /-! From b7666b8884e893f12010025b159a44acf6c293ab Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Pol=27tta=20/=20Miyahara=20K=C5=8D?= Date: Wed, 18 Dec 2024 11:06:48 +0000 Subject: [PATCH 764/829] chore(Mathlib/Tactic/CC/Addition): split a file (#19989) --- Mathlib.lean | 1 + Mathlib/Tactic.lean | 1 + Mathlib/Tactic/CC/Addition.lean | 776 +++---------------------------- Mathlib/Tactic/CC/Datatypes.lean | 7 + Mathlib/Tactic/CC/MkProof.lean | 634 +++++++++++++++++++++++++ scripts/noshake.json | 2 +- 6 files changed, 715 insertions(+), 706 deletions(-) create mode 100644 Mathlib/Tactic/CC/MkProof.lean diff --git a/Mathlib.lean b/Mathlib.lean index 56bf6989e9a40..c7b18fb3042f5 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4720,6 +4720,7 @@ import Mathlib.Tactic.CC import Mathlib.Tactic.CC.Addition import Mathlib.Tactic.CC.Datatypes import Mathlib.Tactic.CC.Lemmas +import Mathlib.Tactic.CC.MkProof import Mathlib.Tactic.CancelDenoms import Mathlib.Tactic.CancelDenoms.Core import Mathlib.Tactic.Cases diff --git a/Mathlib/Tactic.lean b/Mathlib/Tactic.lean index a51290ea7ff52..369301daef0a9 100644 --- a/Mathlib/Tactic.lean +++ b/Mathlib/Tactic.lean @@ -18,6 +18,7 @@ import Mathlib.Tactic.CC import Mathlib.Tactic.CC.Addition import Mathlib.Tactic.CC.Datatypes import Mathlib.Tactic.CC.Lemmas +import Mathlib.Tactic.CC.MkProof import Mathlib.Tactic.CancelDenoms import Mathlib.Tactic.CancelDenoms.Core import Mathlib.Tactic.Cases diff --git a/Mathlib/Tactic/CC/Addition.lean b/Mathlib/Tactic/CC/Addition.lean index 8548c1ad7b0c7..6e7a9fe150ad2 100644 --- a/Mathlib/Tactic/CC/Addition.lean +++ b/Mathlib/Tactic/CC/Addition.lean @@ -3,14 +3,9 @@ Copyright (c) 2016 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Miyahara Kō -/ -import Mathlib.Logic.Basic import Mathlib.Data.Option.Defs import Mathlib.Lean.Expr.Basic -import Mathlib.Tactic.Relation.Rfl -import Mathlib.Tactic.Relation.Symm -import Mathlib.Tactic.CC.Datatypes -import Mathlib.Tactic.CC.Lemmas -import Batteries.Data.RBMap.Alter +import Mathlib.Tactic.CC.MkProof /-! # Process when an new equation is added to a congruence closure @@ -18,25 +13,9 @@ import Batteries.Data.RBMap.Alter universe u -open Lean Meta Elab Tactic Std +open Lean Meta Elab Tactic Std MessageData -namespace Mathlib.Tactic.CC - -initialize - registerTraceClass `Meta.Tactic.cc.merge - registerTraceClass `Meta.Tactic.cc.failure - registerTraceClass `Debug.Meta.Tactic.cc - registerTraceClass `Debug.Meta.Tactic.cc.ac - registerTraceClass `Debug.Meta.Tactic.cc.parentOccs - -/-- The monad for the `cc` tactic stores the current state of the tactic. -/ -abbrev CCM := StateRefT CCStructure MetaM - -namespace CCM - -/-- Run a computation in the `CCM` monad. -/ -@[inline] -def run {α : Type} (x : CCM α) (c : CCStructure) : MetaM (α × CCStructure) := StateRefT'.run x c +namespace Mathlib.Tactic.CC.CCM /-- Update the `todo` field of the state. -/ @[inline] @@ -48,11 +27,6 @@ def modifyTodo (f : Array TodoEntry → Array TodoEntry) : CCM Unit := def modifyACTodo (f : Array ACTodoEntry → Array ACTodoEntry) : CCM Unit := modify fun cc => { cc with acTodo := f cc.acTodo } -/-- Update the `cache` field of the state. -/ -@[inline] -def modifyCache (f : CCCongrTheoremCache → CCCongrTheoremCache) : CCM Unit := - modify fun cc => { cc with cache := f cc.cache } - /-- Read the `todo` field of the state. -/ @[inline] def getTodo : CCM (Array TodoEntry) := do @@ -63,24 +37,6 @@ def getTodo : CCM (Array TodoEntry) := do def getACTodo : CCM (Array ACTodoEntry) := do return (← get).acTodo -/-- Read the `cache` field of the state. -/ -@[inline] -def getCache : CCM CCCongrTheoremCache := do - return (← get).cache - -/-- Look up an entry associated with the given expression. -/ -def getEntry (e : Expr) : CCM (Option Entry) := do - return (← get).entries.find? e - -/-- Use the normalizer to normalize `e`. - -If no normalizer was configured, returns `e` itself. -/ -def normalize (e : Expr) : CCM Expr := do - if let some normalizer := (← get).normalizer then - normalizer.normalize e - else - return e - /-- Add a new entry to the end of the todo list. See also `pushEq`, `pushHEq` and `pushReflEq`. -/ @@ -102,14 +58,6 @@ def pushHEq (lhs rhs : Expr) (H : EntryExpr) : CCM Unit := def pushReflEq (lhs rhs : Expr) : CCM Unit := pushEq lhs rhs .refl -/-- Return the root expression of the expression's congruence class. -/ -def getRoot (e : Expr) : CCM Expr := do - return (← get).root e - -/-- Is `e` the root of its congruence class? -/ -def isCgRoot (e : Expr) : CCM Bool := do - return (← get).isCgRoot e - /-- Update the `child` so its parent becomes `parent`. -/ def addOccurrence (parent child : Expr) (symmTable : Bool) : CCM Unit := do let childRoot ← getRoot child @@ -119,99 +67,6 @@ def addOccurrence (parent child : Expr) (symmTable : Bool) : CCM Unit := do let ps := ps?.getD ∅ ps.insert { expr := parent, symmTable } } -/-- -Return true iff the given function application are congruent - -`e₁` should have the form `f a` and `e₂` the form `g b`. - -See paper: Congruence Closure for Intensional Type Theory. -/ -partial def isCongruent (e₁ e₂ : Expr) : CCM Bool := do - let .app f a := e₁ | failure - let .app g b := e₂ | failure - -- If they are non-dependent functions, then we can compare all arguments at once. - if (← getEntry e₁).any Entry.fo then - e₁.withApp fun f₁ args₁ => - e₂.withApp fun f₂ args₂ => do - if ha : args₁.size = args₂.size then - for hi : i in [:args₁.size] do - if (← getRoot (args₁[i]'hi.2)) != (← getRoot (args₂[i]'(ha.symm ▸ hi.2))) then - return false - if f₁ == f₂ then return true - else if (← getRoot f₁) != (← getRoot f₂) then - -- `f₁` and `f₂` are not equivalent - return false - else if ← pureIsDefEq (← inferType f₁) (← inferType f₂) then - return true - else return false - else return false - else - -- Given `e₁ := f a`, `e₂ := g b` - if (← getRoot a) != (← getRoot b) then - -- `a` and `b` are not equivalent - return false - else if (← getRoot f) != (← getRoot g) then - -- `f` and `g` are not equivalent - return false - else if ← pureIsDefEq (← inferType f) (← inferType g) then - /- Case 1: `f` and `g` have the same type, then we can create a congruence proof for - `HEq (f a) (g b)` -/ - return true - else if f.isApp && g.isApp then - -- Case 2: `f` and `g` are congruent - isCongruent f g - else - /- - f and g are not congruent nor they have the same type. - We can't generate a congruence proof in this case because the following lemma - `hcongr : HEq f₁ f₂ → HEq a₁ a₂ → HEq (f₁ a₁) (f₂ a₂)` - is not provable. - Remark: it is also not provable in MLTT, Coq and Agda (even if we assume UIP). - -/ - return false - -/-- Return the `CongruencesKey` associated with an expression of the form `f a`. -/ -def mkCongruencesKey (e : Expr) : CCM CongruencesKey := do - let .app f a := e | failure - if (← getEntry e).any Entry.fo then - -- first-order case, where we do not consider all partial applications - e.withApp fun fn args => do - return .fo (← getRoot fn) (← args.mapM getRoot) - else - return .ho (← getRoot f) (← getRoot a) - -/-- Return the `SymmCongruencesKey` associated with the equality `lhs = rhs`. -/ -def mkSymmCongruencesKey (lhs rhs : Expr) : CCM SymmCongruencesKey := do - let lhs ← getRoot lhs - let rhs ← getRoot rhs - if hash lhs > hash rhs then return { h₁ := rhs, h₂ := lhs } else return { h₁ := lhs, h₂ := rhs } - -/-- Try to find a congruence theorem for an application of `fn` with `nargs` arguments, with support -for `HEq`. -/ -def mkCCHCongrTheorem (fn : Expr) (nargs : Nat) : CCM (Option CCCongrTheorem) := do - let cache ← getCache - - -- Check if `{ fn, nargs }` is in the cache - let key₁ : CCCongrTheoremKey := { fn, nargs } - if let some it := cache[key₁]? then - return it - - -- Try automatically generated congruence lemma with support for heterogeneous equality. - let lemm ← mkCCHCongrWithArity fn nargs - - if let some lemm := lemm then - modifyCache fun ccc => ccc.insert key₁ (some lemm) - return lemm - - -- cache failure - modifyCache fun ccc => ccc.insert key₁ none - return none - -/-- Try to find a congruence theorem for the expression `e` with support for `HEq`. -/ -def mkCCCongrTheorem (e : Expr) : CCM (Option CCCongrTheorem) := do - let fn := e.getAppFn - let nargs := e.getAppNumArgs - mkCCHCongrTheorem fn nargs - /-- Record the instance `e` and add it to the set of known defeq instances. -/ def propagateInstImplicit (e : Expr) : CCM Unit := do let type ← inferType e @@ -228,399 +83,21 @@ def propagateInstImplicit (e : Expr) : CCM Unit := do modify fun ccs => { ccs with instImplicitReprs := ccs.instImplicitReprs.insert type [e] } -/-- Treat the entry associated with `e` as a first-order function. -/ -def setFO (e : Expr) : CCM Unit := - modify fun ccs => - { ccs with entries := ccs.entries.modify e fun d => { d with fo := true } } - -/-- Update the modification time of the congruence class of `e`. -/ -partial def updateMT (e : Expr) : CCM Unit := do - let r ← getRoot e - let some ps := (← get).parents.find? r | return - for p in ps do - let some it ← getEntry p.expr | failure - let gmt := (← get).gmt - if it.mt < gmt then - let newIt := { it with mt := gmt } - modify fun ccs => - { ccs with entries := ccs.entries.insert p.expr newIt } - updateMT p.expr - -/-- Does the congruence class with root `root` have any `HEq` proofs? -/ -def hasHEqProofs (root : Expr) : CCM Bool := do - let some n ← getEntry root | failure - guard (n.root == root) - return n.heqProofs - -/-- Apply symmetry to `H`, which is an `Eq` or a `HEq`. - -* If `heqProofs` is true, ensure the result is a `HEq` (otherwise it is assumed to be `Eq`). -* If `flipped` is true, apply `symm`, otherwise keep the same direction. -/ -def flipProofCore (H : Expr) (flipped heqProofs : Bool) : CCM Expr := do - let mut newH := H - if ← liftM <| pure heqProofs <&&> Expr.isEq <$> (inferType H >>= whnf) then - newH ← mkAppM ``heq_of_eq #[H] - if !flipped then - return newH - else if heqProofs then - mkHEqSymm newH - else - mkEqSymm newH - -/-- In a delayed way, apply symmetry to `H`, which is an `Eq` or a `HEq`. - -* If `heqProofs` is true, ensure the result is a `HEq` (otherwise it is assumed to be `Eq`). -* If `flipped` is true, apply `symm`, otherwise keep the same direction. -/ -def flipDelayedProofCore (H : DelayedExpr) (flipped heqProofs : Bool) : CCM DelayedExpr := do - let mut newH := H - if heqProofs then - newH := .heqOfEq H - if !flipped then - return newH - else if heqProofs then - return .heqSymm newH - else - return .eqSymm newH - -/-- Apply symmetry to `H`, which is an `Eq` or a `HEq`. - -* If `heqProofs` is true, ensure the result is a `HEq` (otherwise it is assumed to be `Eq`). -* If `flipped` is true, apply `symm`, otherwise keep the same direction. -/ -def flipProof (H : EntryExpr) (flipped heqProofs : Bool) : CCM EntryExpr := - match H with - | .ofExpr H => EntryExpr.ofExpr <$> flipProofCore H flipped heqProofs - | .ofDExpr H => EntryExpr.ofDExpr <$> flipDelayedProofCore H flipped heqProofs - | _ => return H - -/-- Are `e₁` and `e₂` known to be in the same equivalence class? -/ -def isEqv (e₁ e₂ : Expr) : CCM Bool := do - let some n₁ ← getEntry e₁ | return false - let some n₂ ← getEntry e₂ | return false - return n₁.root == n₂.root - -/-- Is `e₁ ≠ e₂` known to be true? - -Note that this is stronger than `not (isEqv e₁ e₂)`: -only if we can prove they are distinct this returns `true`. -/ -def isNotEqv (e₁ e₂ : Expr) : CCM Bool := do - let tmp ← mkEq e₁ e₂ - if ← isEqv tmp (.const ``False []) then return true - let tmp ← mkHEq e₁ e₂ - isEqv tmp (.const ``False []) - -/-- Is the proposition `e` known to be true? -/ -@[inline] -def isEqTrue (e : Expr) : CCM Bool := - isEqv e (.const ``True []) - -/-- Is the proposition `e` known to be false? -/ -@[inline] -def isEqFalse (e : Expr) : CCM Bool := - isEqv e (.const ``False []) - -/-- Apply transitivity to `H₁` and `H₂`, which are both `Eq` or `HEq` depending on `heqProofs`. -/ -def mkTrans (H₁ H₂ : Expr) (heqProofs : Bool) : MetaM Expr := - if heqProofs then mkHEqTrans H₁ H₂ else mkEqTrans H₁ H₂ - -/-- Apply transitivity to `H₁?` and `H₂`, which are both `Eq` or `HEq` depending on `heqProofs`. - -If `H₁?` is `none`, return `H₂` instead. -/ -def mkTransOpt (H₁? : Option Expr) (H₂ : Expr) (heqProofs : Bool) : MetaM Expr := - match H₁? with - | some H₁ => mkTrans H₁ H₂ heqProofs - | none => pure H₂ - -mutual -/-- Use congruence on arguments to prove `lhs = rhs`. - -That is, tries to prove that `lhsFn lhsArgs[0] ... lhsArgs[n-1] = lhsFn rhsArgs[0] ... rhsArgs[n-1]` -by showing that `lhsArgs[i] = rhsArgs[i]` for all `i`. - -Fails if the head function of `lhs` is not that of `rhs`. -/ -partial def mkCongrProofCore (lhs rhs : Expr) (heqProofs : Bool) : CCM Expr := do - let mut lhsArgsRev : Array Expr := #[] - let mut rhsArgsRev : Array Expr := #[] - let mut lhsIt := lhs - let mut rhsIt := rhs - -- Collect the arguments to `lhs` and `rhs`. - -- As an optimization, we stop collecting arguments as soon as the functions are defeq, - -- so `lhsFn` and `rhsFn` might end up still of the form `(f x y z)` and `(f x' y' z')`. - if lhs != rhs then - repeat - let .app lhsItFn lhsItArg := lhsIt | failure - let .app rhsItFn rhsItArg := rhsIt | failure - lhsArgsRev := lhsArgsRev.push lhsItArg - rhsArgsRev := rhsArgsRev.push rhsItArg - lhsIt := lhsItFn - rhsIt := rhsItFn - if lhsIt == rhsIt then - break - if ← pureIsDefEq lhsIt rhsIt then - break - if ← isEqv lhsIt rhsIt <&&> - inferType lhsIt >>= fun i₁ => inferType rhsIt >>= fun i₂ => pureIsDefEq i₁ i₂ then - break - -- If we collect no arguments, the expressions themselves are defeq; return `rfl`. - if lhsArgsRev.isEmpty then - if heqProofs then - return (← mkHEqRefl lhs) - else - return (← mkEqRefl lhs) - let lhsArgs := lhsArgsRev.reverse - let rhsArgs := rhsArgsRev.reverse - -- Ensure that `lhsFn = rhsFn`, they have the same type and the same list of arguments. - let PLift.up ha ← if ha : lhsArgs.size = rhsArgs.size then pure (PLift.up ha) else failure - let lhsFn := lhsIt - let rhsFn := rhsIt - guard (← isEqv lhsFn rhsFn <||> pureIsDefEq lhsFn rhsFn) - guard (← pureIsDefEq (← inferType lhsFn) (← inferType rhsFn)) - /- Create `r`, a proof for - `lhsFn lhsArgs[0] ... lhsArgs[n-1] = lhsFn rhsArgs[0] ... rhsArgs[n-1]` - where `n := lhsArgs.size` -/ - let some specLemma ← mkCCHCongrTheorem lhsFn lhsArgs.size | failure - let mut kindsIt := specLemma.argKinds - let mut lemmaArgs : Array Expr := #[] - for hi : i in [:lhsArgs.size] do - guard !kindsIt.isEmpty - lemmaArgs := lemmaArgs.push (lhsArgs[i]'hi.2) |>.push (rhsArgs[i]'(ha.symm ▸ hi.2)) - if kindsIt[0]! matches CongrArgKind.heq then - let some p ← getHEqProof (lhsArgs[i]'hi.2) (rhsArgs[i]'(ha.symm ▸ hi.2)) | failure - lemmaArgs := lemmaArgs.push p - else - guard (kindsIt[0]! matches .eq) - let some p ← getEqProof (lhsArgs[i]'hi.2) (rhsArgs[i]'(ha.symm ▸ hi.2)) | failure - lemmaArgs := lemmaArgs.push p - kindsIt := kindsIt.eraseIdx! 0 - let mut r := mkAppN specLemma.proof lemmaArgs - if specLemma.heqResult && !heqProofs then - r ← mkAppM ``eq_of_heq #[r] - else if !specLemma.heqResult && heqProofs then - r ← mkAppM ``heq_of_eq #[r] - if ← pureIsDefEq lhsFn rhsFn then - return r - /- Convert `r` into a proof of `lhs = rhs` using `Eq.rec` and - the proof that `lhsFn = rhsFn` -/ - let some lhsFnEqRhsFn ← getEqProof lhsFn rhsFn | failure - let motive ← - withLocalDeclD `x (← inferType lhsFn) fun x => do - let motiveRhs := mkAppN x rhsArgs - let motive ← if heqProofs then mkHEq lhs motiveRhs else mkEq lhs motiveRhs - let hType ← mkEq lhsFn x - withLocalDeclD `h hType fun h => - mkLambdaFVars #[x, h] motive - mkEqRec motive r lhsFnEqRhsFn - -/-- If `e₁ : R lhs₁ rhs₁`, `e₂ : R lhs₂ rhs₂` and `lhs₁ = rhs₂`, where `R` is a symmetric relation, -prove `R lhs₁ rhs₁` is equivalent to `R lhs₂ rhs₂`. - - * if `lhs₁` is known to equal `lhs₂`, return `none` - * if `lhs₁` is not known to equal `rhs₂`, fail. -/ -partial def mkSymmCongrProof (e₁ e₂ : Expr) (heqProofs : Bool) : CCM (Option Expr) := do - let some (R₁, lhs₁, rhs₁) ← e₁.relSidesIfSymm? | return none - let some (R₂, lhs₂, rhs₂) ← e₂.relSidesIfSymm? | return none - if R₁ != R₂ then return none - if (← isEqv lhs₁ lhs₂) then - return none - guard (← isEqv lhs₁ rhs₂) - /- We must apply symmetry. - The symm congruence table is implicitly using symmetry. - That is, we have - `e₁ := lhs₁ ~R₁~ rhs₁` - and - `e2 := lhs₂ ~R₁~ rhs₂` - But, - `lhs₁ ~R₁~ rhs₂` and `rhs₁ ~R₁~ lhs₂` -/ - /- Given `e₁ := lhs₁ ~R₁~ rhs₁`, - create proof for - `lhs₁ ~R₁~ rhs₁` = `rhs₁ ~R₁~ lhs₁` -/ - let newE₁ ← mkRel R₁ rhs₁ lhs₁ - let e₁IffNewE₁ ← - withLocalDeclD `h₁ e₁ fun h₁ => - withLocalDeclD `h₂ newE₁ fun h₂ => do - mkAppM ``Iff.intro - #[← mkLambdaFVars #[h₁] (← h₁.applySymm), ← mkLambdaFVars #[h₂] (← h₂.applySymm)] - let mut e₁EqNewE₁ := mkApp3 (.const ``propext []) e₁ newE₁ e₁IffNewE₁ - let newE₁EqE₂ ← mkCongrProofCore newE₁ e₂ heqProofs - if heqProofs then - e₁EqNewE₁ ← mkAppM ``heq_of_eq #[e₁EqNewE₁] - return some (← mkTrans e₁EqNewE₁ newE₁EqE₂ heqProofs) - -/-- Use congruence on arguments to prove `e₁ = e₂`. - -Special case: if `e₁` and `e₂` have the form `R lhs₁ rhs₁` and `R lhs₂ rhs₂` such that -`R` is symmetric and `lhs₁ = rhs₂`, then use those facts instead. -/ -partial def mkCongrProof (e₁ e₂ : Expr) (heqProofs : Bool) : CCM Expr := do - if let some r ← mkSymmCongrProof e₁ e₂ heqProofs then - return r +/-- Return the `CongruencesKey` associated with an expression of the form `f a`. -/ +def mkCongruencesKey (e : Expr) : CCM CongruencesKey := do + let .app f a := e | failure + if (← getEntry e).any Entry.fo then + -- first-order case, where we do not consider all partial applications + e.withApp fun fn args => do + return .fo (← getRoot fn) (← args.mapM getRoot) else - mkCongrProofCore e₁ e₂ heqProofs - -/-- Turn a delayed proof into an actual proof term. -/ -partial def mkDelayedProof (H : DelayedExpr) : CCM Expr := do - match H with - | .ofExpr H => return H - | .eqProof lhs rhs => liftOption (← getEqProof lhs rhs) - | .congrArg f h => mkCongrArg f (← mkDelayedProof h) - | .congrFun h a => mkCongrFun (← mkDelayedProof h) (← liftOption a.toExpr) - | .eqSymm h => mkEqSymm (← mkDelayedProof h) - | .eqSymmOpt a₁ a₂ h => - mkAppOptM ``Eq.symm #[none, ← liftOption a₁.toExpr, ← liftOption a₂.toExpr, ← mkDelayedProof h] - | .eqTrans h₁ h₂ => mkEqTrans (← mkDelayedProof h₁) (← mkDelayedProof h₂) - | .eqTransOpt a₁ a₂ a₃ h₁ h₂ => - mkAppOptM ``Eq.trans - #[none, ← liftOption a₁.toExpr, ← liftOption a₂.toExpr, ← liftOption a₃.toExpr, - ← mkDelayedProof h₁, ← mkDelayedProof h₂] - | .heqOfEq h => mkAppM ``heq_of_eq #[← mkDelayedProof h] - | .heqSymm h => mkHEqSymm (← mkDelayedProof h) - -/-- Use the format of `H` to try and construct a proof or `lhs = rhs`: - * If `H = .congr`, then use congruence. - * If `H = .eqTrue`, try to prove `lhs = True` or `rhs = True`, - if they have the format `R a b`, by proving `a = b`. - * Otherwise, return the (delayed) proof encoded by `H` itself. -/ -partial def mkProof (lhs rhs : Expr) (H : EntryExpr) (heqProofs : Bool) : CCM Expr := do - match H with - | .congr => mkCongrProof lhs rhs heqProofs - | .eqTrue => - let (flip, some (R, a, b)) ← - if lhs == .const ``True [] then - ((true, ·)) <$> rhs.relSidesIfRefl? - else - ((false, ·)) <$> lhs.relSidesIfRefl? - | failure - let aRb ← - if R == ``Eq then - getEqProof a b >>= liftOption - else if R == ``HEq then - getHEqProof a b >>= liftOption - else - -- TODO(Leo): the following code assumes R is homogeneous. - -- We should add support arbitrary heterogeneous reflexive relations. - getEqProof a b >>= liftOption >>= fun aEqb => liftM (liftFromEq R aEqb) - let aRbEqTrue ← mkEqTrue aRb - if flip then - mkEqSymm aRbEqTrue - else - return aRbEqTrue - | .refl => - let type ← if heqProofs then mkHEq lhs rhs else mkEq lhs rhs - let proof ← if heqProofs then mkHEqRefl lhs else mkEqRefl lhs - mkExpectedTypeHint proof type - | .ofExpr H => return H - | .ofDExpr H => mkDelayedProof H - -/-- -If `asHEq` is `true`, then build a proof for `HEq e₁ e₂`. -Otherwise, build a proof for `e₁ = e₂`. -The result is `none` if `e₁` and `e₂` are not in the same equivalence class. -/ -partial def getEqProofCore (e₁ e₂ : Expr) (asHEq : Bool) : CCM (Option Expr) := do - if e₁.hasExprMVar || e₂.hasExprMVar then return none - if ← pureIsDefEq e₁ e₂ then - if asHEq then - return some (← mkHEqRefl e₁) - else - return some (← mkEqRefl e₁) - let some n₁ ← getEntry e₁ | return none - let some n₂ ← getEntry e₂ | return none - if n₁.root != n₂.root then return none - let heqProofs ← hasHEqProofs n₁.root - -- 1. Retrieve "path" from `e₁` to `root` - let mut path₁ : Array Expr := #[] - let mut Hs₁ : Array EntryExpr := #[] - let mut visited : RBExprSet := ∅ - let mut it₁ := e₁ - repeat - visited := visited.insert it₁ - let some it₁N ← getEntry it₁ | failure - let some t := it₁N.target | break - path₁ := path₁.push t - let some p := it₁N.proof | failure - Hs₁ := Hs₁.push (← flipProof p it₁N.flipped heqProofs) - it₁ := t - guard (it₁ == n₁.root) - -- 2. The path from `e₂` to root must have at least one element `c` in visited - -- Retrieve "path" from `e₂` to `c` - let mut path₂ : Array Expr := #[] - let mut Hs₂ : Array EntryExpr := #[] - let mut it₂ := e₂ - repeat - if visited.contains it₂ then - break -- found common - let some it₂N ← getEntry it₂ | failure - let some t := it₂N.target | failure - path₂ := path₂.push it₂ - let some p := it₂N.proof | failure - Hs₂ := Hs₂.push (← flipProof p (!it₂N.flipped) heqProofs) - it₂ := t - -- `it₂` is the common element... - -- 3. Shrink `path₁`/`Hs₁` until we find `it₂` (the common element) - repeat - if path₁.isEmpty then - guard (it₂ == e₁) - break - if path₁.back! == it₂ then - -- found it! - break - path₁ := path₁.pop - Hs₁ := Hs₁.pop - - -- 4. Build transitivity proof - let mut pr? : Option Expr := none - let mut lhs := e₁ - for i in [:path₁.size] do - pr? ← some <$> mkTransOpt pr? (← mkProof lhs path₁[i]! Hs₁[i]! heqProofs) heqProofs - lhs := path₁[i]! - let mut i := Hs₂.size - while i > 0 do - i := i - 1 - pr? ← some <$> mkTransOpt pr? (← mkProof lhs path₂[i]! Hs₂[i]! heqProofs) heqProofs - lhs := path₂[i]! - let mut some pr := pr? | failure - if heqProofs && !asHEq then - pr ← mkAppM ``eq_of_heq #[pr] - else if !heqProofs && asHEq then - pr ← mkAppM ``heq_of_eq #[pr] - return pr - -/-- Build a proof for `e₁ = e₂`. -The result is `none` if `e₁` and `e₂` are not in the same equivalence class. -/ -@[inline] -partial def getEqProof (e₁ e₂ : Expr) : CCM (Option Expr) := - getEqProofCore e₁ e₂ false - -/-- Build a proof for `HEq e₁ e₂`. -The result is `none` if `e₁` and `e₂` are not in the same equivalence class. -/ -@[inline] -partial def getHEqProof (e₁ e₂ : Expr) : CCM (Option Expr) := - getEqProofCore e₁ e₂ true -end + return .ho (← getRoot f) (← getRoot a) -/-- Build a proof for `e = True`. Fails if `e` is not known to be true. -/ -def getEqTrueProof (e : Expr) : CCM Expr := do - guard (← isEqTrue e) - let some p ← getEqProof e (.const ``True []) | failure - return p - -/-- Build a proof for `e = False`. Fails if `e` is not known to be false. -/ -def getEqFalseProof (e : Expr) : CCM Expr := do - guard (← isEqFalse e) - let some p ← getEqProof e (.const ``False []) | failure - return p - -/-- Build a proof for `a = b`. Fails if `a` and `b` are not known to be equal. -/ -def getPropEqProof (a b : Expr) : CCM Expr := do - guard (← isEqv a b) - let some p ← getEqProof a b | failure - return p - -/-- Build a proof of `False` if the context is inconsistent. -Returns `none` if `False` is not known to be true. -/ -def getInconsistencyProof : CCM (Option Expr) := do - guard !(← get).frozePartitions - if let some p ← getEqProof (.const ``True []) (.const ``False []) then - return some (← mkAppM ``false_of_true_eq_false #[p]) - else - return none +/-- Return the `SymmCongruencesKey` associated with the equality `lhs = rhs`. -/ +def mkSymmCongruencesKey (lhs rhs : Expr) : CCM SymmCongruencesKey := do + let lhs ← getRoot lhs + let rhs ← getRoot rhs + if hash lhs > hash rhs then return { h₁ := rhs, h₂ := lhs } else return { h₁ := lhs, h₂ := rhs } /-- Auxiliary function for comparing `lhs₁ ~ rhs₁` and `lhs₂ ~ rhs₂`, when `~` is symmetric/commutative. @@ -774,57 +251,6 @@ def propagateBeta (fn : Expr) (revArgs : Array Expr) (lambdas : Array Expr) newLambdaApps := newLambdaApps.push newApp return newLambdaApps -/-- Given `a`, `a₁` and `a₁NeB : a₁ ≠ b`, return a proof of `a ≠ b` if `a` and `a₁` are in the -same equivalence class. -/ -def mkNeOfEqOfNe (a a₁ a₁NeB : Expr) : CCM (Option Expr) := do - guard (← isEqv a a₁) - if a == a₁ then - return some a₁NeB - let aEqA₁ ← getEqProof a a₁ - match aEqA₁ with - | none => return none -- failed to build proof - | some aEqA₁ => mkAppM ``ne_of_eq_of_ne #[aEqA₁, a₁NeB] - -/-- Given `aNeB₁ : a ≠ b₁`, `b₁` and `b`, return a proof of `a ≠ b` if `b` and `b₁` are in the -same equivalence class. -/ -def mkNeOfNeOfEq (aNeB₁ b₁ b : Expr) : CCM (Option Expr) := do - guard (← isEqv b b₁) - if b == b₁ then - return some aNeB₁ - let b₁EqB ← getEqProof b b₁ - match b₁EqB with - | none => return none -- failed to build proof - | some b₁EqB => mkAppM ``ne_of_ne_of_eq #[aNeB₁, b₁EqB] - -/-- If `e` is of the form `op e₁ e₂` where `op` is an associative and commutative binary operator, -return the canonical form of `op`. -/ -def isAC (e : Expr) : CCM (Option Expr) := do - let .app (.app op _) _ := e | return none - let ccs ← get - if let some cop := ccs.canOps.find? op then - let some b := ccs.opInfo.find? cop - | throwError "opInfo should contain all canonical operators in canOps" - return bif b then some cop else none - for (cop, b) in ccs.opInfo do - if ← pureIsDefEq op cop then - modify fun _ => { ccs with canOps := ccs.canOps.insert op cop } - return bif b then some cop else none - let b ← - try - let aop ← mkAppM ``Std.Associative #[op] - let some _ ← synthInstance? aop | failure - let cop ← mkAppM ``Std.Commutative #[op] - let some _ ← synthInstance? cop | failure - pure true - catch _ => - pure false - modify fun _ => - { ccs with - canOps := ccs.canOps.insert op op - opInfo := ccs.opInfo.insert op b } - return bif b then some op else none - -open MessageData in /-- Given `lhs`, `rhs`, and `header := "my header:"`, Trace `my header: lhs = rhs`. -/ def dbgTraceACEq (header : String) (lhs rhs : ACApps) : CCM Unit := do let ccs ← get @@ -832,119 +258,11 @@ def dbgTraceACEq (header : String) (lhs rhs : ACApps) : CCM Unit := do group (ofFormat (header ++ .line) ++ ccs.ppACApps lhs ++ ofFormat (.line ++ "=" ++ .line) ++ ccs.ppACApps rhs) -open MessageData in /-- Trace the state of AC module. -/ def dbgTraceACState : CCM Unit := do let ccs ← get trace[Debug.Meta.Tactic.cc.ac] group ("state: " ++ nest 6 ccs.ppAC) -/-- Return the proof of `e₁ = e₂` using `ac_rfl` tactic. -/ -def mkACProof (e₁ e₂ : Expr) : MetaM Expr := do - let eq ← mkEq e₁ e₂ - let .mvar m ← mkFreshExprSyntheticOpaqueMVar eq | failure - AC.rewriteUnnormalizedRefl m - let pr ← instantiateMVars (.mvar m) - mkExpectedTypeHint pr eq - -/-- Given `tr := t*r` `sr := s*r` `tEqs : t = s`, return a proof for `tr = sr` - - We use `a*b` to denote an AC application. That is, `(a*b)*(c*a)` is the term `a*a*b*c`. -/ -def mkACSimpProof (tr t s r sr : ACApps) (tEqs : DelayedExpr) : MetaM DelayedExpr := do - if tr == t then - return tEqs - else if tr == sr then - let some tre := tr.toExpr | failure - DelayedExpr.ofExpr <$> mkEqRefl tre - else - let .apps op _ := tr | failure - let some re := r.toExpr | failure - let some te := t.toExpr | failure - let some se := s.toExpr | failure - let some tre := tr.toExpr | failure - let some sre := sr.toExpr | failure - let opr := op.app re -- `(*) r` - let rt := mkApp2 op re te -- `r * t` - let rs := mkApp2 op re se -- `r * s` - let rtEqrs := DelayedExpr.congrArg opr tEqs - let trEqrt ← mkACProof tre rt - let rsEqsr ← mkACProof rs sre - return .eqTrans (.eqTrans trEqrt rtEqrs) rsEqsr - -/-- Given `ra := a*r` `sb := b*s` `ts := t*s` `tr := t*r` `tsEqa : t*s = a` `trEqb : t*r = b`, - return a proof for `ra = sb`. - - We use `a*b` to denote an AC application. That is, `(a*b)*(c*a)` is the term `a*a*b*c`. -/ -def mkACSuperposeProof (ra sb a b r s ts tr : ACApps) (tsEqa trEqb : DelayedExpr) : - MetaM DelayedExpr := do - let .apps _ _ := tr | failure - let .apps op _ := ts | failure - let some tse := ts.toExpr | failure - let some re := r.toExpr | failure - let some tre := tr.toExpr | failure - let some se := s.toExpr | failure - let some ae := a.toExpr | failure - let some be := b.toExpr | failure - let some rae := ra.toExpr | failure - let some sbe := sb.toExpr | failure - let tsrEqar := DelayedExpr.congrFun (.congrArg op tsEqa) r -- `(t * s) * r = a * r` - let trsEqbs := DelayedExpr.congrFun (.congrArg op trEqb) s -- `(t * r) * s = b * s` - let tsr := mkApp2 op tse re -- `(t * s) * r` - let trs := mkApp2 op tre se -- `(t * r) * s` - let ar := mkApp2 op ae re -- `a * r` - let bs := mkApp2 op be se -- `b * r` - let tsrEqtrs ← mkACProof tsr trs -- `(t * s) * r = (t * r) * s` - let raEqar ← mkACProof rae ar -- `r * a = a * r` - let bsEqsb ← mkACProof bs sbe -- `b * s = s * b` - return .eqTrans raEqar (.eqTrans (.eqSymm tsrEqar) (.eqTrans tsrEqtrs (.eqTrans trsEqbs bsEqsb))) - -/-- Given `e := lhs * r` and `H : lhs = rhs`, return `rhs * r` and the proof of `e = rhs * r`. -/ -def simplifyACCore (e lhs rhs : ACApps) (H : DelayedExpr) : - CCM (ACApps × DelayedExpr) := do - guard (lhs.isSubset e) - if e == lhs then - return (rhs, H) - else - let .apps op _ := e | failure - let newArgs := e.diff lhs - let r : ACApps := if newArgs.isEmpty then default else .mkApps op newArgs - let newArgs := ACApps.append op rhs newArgs - let newE := ACApps.mkApps op newArgs - let some true := (← get).opInfo.find? op | failure - let newPr ← mkACSimpProof e lhs rhs r newE H - return (newE, newPr) - -/-- The single step of `simplifyAC`. - -Simplifies an expression `e` by either simplifying one argument to the AC operator, or the whole -expression. -/ -def simplifyACStep (e : ACApps) : CCM (Option (ACApps × DelayedExpr)) := do - if let .apps _ args := e then - for h : i in [:args.size] do - if i == 0 || (args[i]'h.2) != (args[i - 1]'(Nat.lt_of_le_of_lt (i.sub_le 1) h.2)) then - let some ae := (← get).acEntries.find? (args[i]'h.2) | failure - let occs := ae.RLHSOccs - let mut Rlhs? : Option ACApps := none - for Rlhs in occs do - if Rlhs.isSubset e then - Rlhs? := some Rlhs - break - if let some Rlhs := Rlhs? then - let some (Rrhs, H) := (← get).acR.find? Rlhs | failure - return (some <| ← simplifyACCore e Rlhs Rrhs H) - else if let some p := (← get).acR.find? e then - return some p - return none - -/-- If `e` can be simplified by the AC module, return the simplified term and the proof term of the -equality. -/ -def simplifyAC (e : ACApps) : CCM (Option (ACApps × DelayedExpr)) := do - let mut some (curr, pr) ← simplifyACStep e | return none - repeat - let some (newCurr, newPr) ← simplifyACStep curr | break - pr := .eqTransOpt e curr newCurr pr newPr - curr := newCurr - return some (curr, pr) - /-- Insert or erase `lhs` to the occurrences of `arg` on an equality in `acR`. -/ def insertEraseROcc (arg : Expr) (lhs : ACApps) (inLHS isInsert : Bool) : CCM Unit := do let some entry := (← get).acEntries.find? arg | failure @@ -1000,7 +318,6 @@ an equality in `acR`. -/ def eraseRRHSOccs (e lhs : ACApps) : CCM Unit := eraseROccs e lhs false -open MessageData in /-- Try to simplify the right hand sides of equalities in `acR` by `H : lhs = rhs`. -/ def composeAC (lhs rhs : ACApps) (H : DelayedExpr) : CCM Unit := do let some x := (← get).getVarWithLeastRHSOccs lhs | failure @@ -1024,7 +341,6 @@ def composeAC (lhs rhs : ACApps) (H : DelayedExpr) : CCM Unit := do (oldRw ++ ofFormat (Format.line ++ "with" ++ .line) ++ newRw) ++ ofFormat (Format.line ++ ":=" ++ .line) ++ ccs.ppACApps newRrhs) -open MessageData in /-- Try to simplify the left hand sides of equalities in `acR` by `H : lhs = rhs`. -/ def collapseAC (lhs rhs : ACApps) (H : DelayedExpr) : CCM Unit := do let some x := (← get).getVarWithLeastLHSOccs lhs | failure @@ -1049,7 +365,33 @@ def collapseAC (lhs rhs : ACApps) (H : DelayedExpr) : CCM Unit := do (newRw ++ ofFormat (Format.line ++ "at" ++ .line) ++ oldRw) ++ ofFormat (Format.line ++ ":=" ++ .line) ++ ccs.ppACApps newRlhs) -open MessageData in +/-- Given `ra := a*r` `sb := b*s` `ts := t*s` `tr := t*r` `tsEqa : t*s = a` `trEqb : t*r = b`, + return a proof for `ra = sb`. + + We use `a*b` to denote an AC application. That is, `(a*b)*(c*a)` is the term `a*a*b*c`. -/ +def mkACSuperposeProof (ra sb a b r s ts tr : ACApps) (tsEqa trEqb : DelayedExpr) : + MetaM DelayedExpr := do + let .apps _ _ := tr | failure + let .apps op _ := ts | failure + let some tse := ts.toExpr | failure + let some re := r.toExpr | failure + let some tre := tr.toExpr | failure + let some se := s.toExpr | failure + let some ae := a.toExpr | failure + let some be := b.toExpr | failure + let some rae := ra.toExpr | failure + let some sbe := sb.toExpr | failure + let tsrEqar := DelayedExpr.congrFun (.congrArg op tsEqa) r -- `(t * s) * r = a * r` + let trsEqbs := DelayedExpr.congrFun (.congrArg op trEqb) s -- `(t * r) * s = b * s` + let tsr := mkApp2 op tse re -- `(t * s) * r` + let trs := mkApp2 op tre se -- `(t * r) * s` + let ar := mkApp2 op ae re -- `a * r` + let bs := mkApp2 op be se -- `b * r` + let tsrEqtrs ← mkACProof tsr trs -- `(t * s) * r = (t * r) * s` + let raEqar ← mkACProof rae ar -- `r * a = a * r` + let bsEqsb ← mkACProof bs sbe -- `b * s = s * b` + return .eqTrans raEqar (.eqTrans (.eqSymm tsrEqar) (.eqTrans tsrEqtrs (.eqTrans trsEqbs bsEqsb))) + /-- Given `tsEqa : ts = a`, for each equality `trEqb : tr = b` in `acR` where the intersection `t` of `ts` and `tr` is nonempty, let `ts = t*s` and `tr := t*r`, add a new equality `r*a = s*b`. -/ @@ -1089,7 +431,6 @@ def superposeAC (ts a : ACApps) (tsEqa : DelayedExpr) : CCM Unit := do (rw₁ ++ ofFormat (Format.line ++ "with" ++ .line) ++ rw₂) ++ ofFormat (Format.line ++ ":=" ++ .line) ++ eq) -open MessageData in /-- Process the tasks in the `acTodo` field. -/ def processAC : CCM Unit := do repeat @@ -1172,6 +513,34 @@ def internalizeACVar (e : Expr) : CCM Bool := do setACVar e return true +/-- If `e` is of the form `op e₁ e₂` where `op` is an associative and commutative binary operator, +return the canonical form of `op`. -/ +def isAC (e : Expr) : CCM (Option Expr) := do + let .app (.app op _) _ := e | return none + let ccs ← get + if let some cop := ccs.canOps.find? op then + let some b := ccs.opInfo.find? cop + | throwError "opInfo should contain all canonical operators in canOps" + return bif b then some cop else none + for (cop, b) in ccs.opInfo do + if ← pureIsDefEq op cop then + modify fun _ => { ccs with canOps := ccs.canOps.insert op cop } + return bif b then some cop else none + let b ← + try + let aop ← mkAppM ``Std.Associative #[op] + let some _ ← synthInstance? aop | failure + let cop ← mkAppM ``Std.Commutative #[op] + let some _ ← synthInstance? cop | failure + pure true + catch _ => + pure false + modify fun _ => + { ccs with + canOps := ccs.canOps.insert op op + opInfo := ccs.opInfo.insert op b } + return bif b then some op else none + /-- Given `e := op₁ (op₂ a₁ a₂) (op₃ a₃ a₄)` where `opₙ`s are canonicalized to `op`, internalize `aₙ`s as AC variables and return `(op (op a₁ a₂) (op a₃ a₄), args ++ #[a₁, a₂, a₃, a₄])`. -/ partial def convertAC (op e : Expr) (args : Array Expr := #[]) : CCM (Array Expr × Expr) := do @@ -1184,7 +553,6 @@ partial def convertAC (op e : Expr) (args : Array Expr := #[]) : CCM (Array Expr let _ ← internalizeACVar e return (args.push e, e) -open MessageData in /-- Internalize `e` so that the AC module can deal with the given expression. If the expression does not contain an AC operator, or the parent expression @@ -2101,5 +1469,3 @@ def add (type : Expr) (proof : Expr) : CCM Unit := do end CCM end Mathlib.Tactic.CC - -set_option linter.style.longFile 2300 diff --git a/Mathlib/Tactic/CC/Datatypes.lean b/Mathlib/Tactic/CC/Datatypes.lean index 6c11094d2ff5a..32239f6b361a3 100644 --- a/Mathlib/Tactic/CC/Datatypes.lean +++ b/Mathlib/Tactic/CC/Datatypes.lean @@ -692,4 +692,11 @@ structure CCStructure extends CCState where cache : CCCongrTheoremCache := ∅ deriving Inhabited +initialize + registerTraceClass `Meta.Tactic.cc.merge + registerTraceClass `Meta.Tactic.cc.failure + registerTraceClass `Debug.Meta.Tactic.cc + registerTraceClass `Debug.Meta.Tactic.cc.ac + registerTraceClass `Debug.Meta.Tactic.cc.parentOccs + end Mathlib.Tactic.CC diff --git a/Mathlib/Tactic/CC/MkProof.lean b/Mathlib/Tactic/CC/MkProof.lean new file mode 100644 index 0000000000000..c3827b0526b16 --- /dev/null +++ b/Mathlib/Tactic/CC/MkProof.lean @@ -0,0 +1,634 @@ +/- +Copyright (c) 2016 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Miyahara Kō +-/ +import Batteries.Data.RBMap.Alter +import Mathlib.Logic.Basic +import Mathlib.Tactic.CC.Datatypes +import Mathlib.Tactic.CC.Lemmas +import Mathlib.Tactic.Relation.Rfl +import Mathlib.Tactic.Relation.Symm + +/-! +# Make proofs from a congruence closure +-/ + +open Lean Meta Elab Tactic Std + +namespace Mathlib.Tactic.CC + +/-- The monad for the `cc` tactic stores the current state of the tactic. -/ +abbrev CCM := StateRefT CCStructure MetaM + +namespace CCM + +/-- Run a computation in the `CCM` monad. -/ +@[inline] +def run {α : Type} (x : CCM α) (c : CCStructure) : MetaM (α × CCStructure) := StateRefT'.run x c + +/-- Update the `cache` field of the state. -/ +@[inline] +def modifyCache (f : CCCongrTheoremCache → CCCongrTheoremCache) : CCM Unit := + modify fun cc => { cc with cache := f cc.cache } + +/-- Read the `cache` field of the state. -/ +@[inline] +def getCache : CCM CCCongrTheoremCache := do + return (← get).cache + +/-- Look up an entry associated with the given expression. -/ +def getEntry (e : Expr) : CCM (Option Entry) := do + return (← get).entries.find? e + +/-- Use the normalizer to normalize `e`. + +If no normalizer was configured, returns `e` itself. -/ +def normalize (e : Expr) : CCM Expr := do + if let some normalizer := (← get).normalizer then + normalizer.normalize e + else + return e + +/-- Return the root expression of the expression's congruence class. -/ +def getRoot (e : Expr) : CCM Expr := do + return (← get).root e + +/-- Is `e` the root of its congruence class? -/ +def isCgRoot (e : Expr) : CCM Bool := do + return (← get).isCgRoot e + +/-- +Return true iff the given function application are congruent + +`e₁` should have the form `f a` and `e₂` the form `g b`. + +See paper: Congruence Closure for Intensional Type Theory. -/ +partial def isCongruent (e₁ e₂ : Expr) : CCM Bool := do + let .app f a := e₁ | failure + let .app g b := e₂ | failure + -- If they are non-dependent functions, then we can compare all arguments at once. + if (← getEntry e₁).any Entry.fo then + e₁.withApp fun f₁ args₁ => + e₂.withApp fun f₂ args₂ => do + if ha : args₁.size = args₂.size then + for hi : i in [:args₁.size] do + if (← getRoot (args₁[i]'hi.2)) != (← getRoot (args₂[i]'(ha.symm ▸ hi.2))) then + return false + if f₁ == f₂ then return true + else if (← getRoot f₁) != (← getRoot f₂) then + -- `f₁` and `f₂` are not equivalent + return false + else if ← pureIsDefEq (← inferType f₁) (← inferType f₂) then + return true + else return false + else return false + else + -- Given `e₁ := f a`, `e₂ := g b` + if (← getRoot a) != (← getRoot b) then + -- `a` and `b` are not equivalent + return false + else if (← getRoot f) != (← getRoot g) then + -- `f` and `g` are not equivalent + return false + else if ← pureIsDefEq (← inferType f) (← inferType g) then + /- Case 1: `f` and `g` have the same type, then we can create a congruence proof for + `HEq (f a) (g b)` -/ + return true + else if f.isApp && g.isApp then + -- Case 2: `f` and `g` are congruent + isCongruent f g + else + /- + f and g are not congruent nor they have the same type. + We can't generate a congruence proof in this case because the following lemma + `hcongr : HEq f₁ f₂ → HEq a₁ a₂ → HEq (f₁ a₁) (f₂ a₂)` + is not provable. + Remark: it is also not provable in MLTT, Coq and Agda (even if we assume UIP). + -/ + return false + +/-- Try to find a congruence theorem for an application of `fn` with `nargs` arguments, with support +for `HEq`. -/ +def mkCCHCongrTheorem (fn : Expr) (nargs : Nat) : CCM (Option CCCongrTheorem) := do + let cache ← getCache + + -- Check if `{ fn, nargs }` is in the cache + let key₁ : CCCongrTheoremKey := { fn, nargs } + if let some it := cache[key₁]? then + return it + + -- Try automatically generated congruence lemma with support for heterogeneous equality. + let lemm ← mkCCHCongrWithArity fn nargs + + if let some lemm := lemm then + modifyCache fun ccc => ccc.insert key₁ (some lemm) + return lemm + + -- cache failure + modifyCache fun ccc => ccc.insert key₁ none + return none + +/-- Try to find a congruence theorem for the expression `e` with support for `HEq`. -/ +def mkCCCongrTheorem (e : Expr) : CCM (Option CCCongrTheorem) := do + let fn := e.getAppFn + let nargs := e.getAppNumArgs + mkCCHCongrTheorem fn nargs + +/-- Treat the entry associated with `e` as a first-order function. -/ +def setFO (e : Expr) : CCM Unit := + modify fun ccs => + { ccs with entries := ccs.entries.modify e fun d => { d with fo := true } } + +/-- Update the modification time of the congruence class of `e`. -/ +partial def updateMT (e : Expr) : CCM Unit := do + let r ← getRoot e + let some ps := (← get).parents.find? r | return + for p in ps do + let some it ← getEntry p.expr | failure + let gmt := (← get).gmt + if it.mt < gmt then + let newIt := { it with mt := gmt } + modify fun ccs => + { ccs with entries := ccs.entries.insert p.expr newIt } + updateMT p.expr + +/-- Does the congruence class with root `root` have any `HEq` proofs? -/ +def hasHEqProofs (root : Expr) : CCM Bool := do + let some n ← getEntry root | failure + guard (n.root == root) + return n.heqProofs + +/-- Apply symmetry to `H`, which is an `Eq` or a `HEq`. + +* If `heqProofs` is true, ensure the result is a `HEq` (otherwise it is assumed to be `Eq`). +* If `flipped` is true, apply `symm`, otherwise keep the same direction. -/ +def flipProofCore (H : Expr) (flipped heqProofs : Bool) : CCM Expr := do + let mut newH := H + if ← liftM <| pure heqProofs <&&> Expr.isEq <$> (inferType H >>= whnf) then + newH ← mkAppM ``heq_of_eq #[H] + if !flipped then + return newH + else if heqProofs then + mkHEqSymm newH + else + mkEqSymm newH + +/-- In a delayed way, apply symmetry to `H`, which is an `Eq` or a `HEq`. + +* If `heqProofs` is true, ensure the result is a `HEq` (otherwise it is assumed to be `Eq`). +* If `flipped` is true, apply `symm`, otherwise keep the same direction. -/ +def flipDelayedProofCore (H : DelayedExpr) (flipped heqProofs : Bool) : CCM DelayedExpr := do + let mut newH := H + if heqProofs then + newH := .heqOfEq H + if !flipped then + return newH + else if heqProofs then + return .heqSymm newH + else + return .eqSymm newH + +/-- Apply symmetry to `H`, which is an `Eq` or a `HEq`. + +* If `heqProofs` is true, ensure the result is a `HEq` (otherwise it is assumed to be `Eq`). +* If `flipped` is true, apply `symm`, otherwise keep the same direction. -/ +def flipProof (H : EntryExpr) (flipped heqProofs : Bool) : CCM EntryExpr := + match H with + | .ofExpr H => EntryExpr.ofExpr <$> flipProofCore H flipped heqProofs + | .ofDExpr H => EntryExpr.ofDExpr <$> flipDelayedProofCore H flipped heqProofs + | _ => return H + +/-- Are `e₁` and `e₂` known to be in the same equivalence class? -/ +def isEqv (e₁ e₂ : Expr) : CCM Bool := do + let some n₁ ← getEntry e₁ | return false + let some n₂ ← getEntry e₂ | return false + return n₁.root == n₂.root + +/-- Is `e₁ ≠ e₂` known to be true? + +Note that this is stronger than `not (isEqv e₁ e₂)`: +only if we can prove they are distinct this returns `true`. -/ +def isNotEqv (e₁ e₂ : Expr) : CCM Bool := do + let tmp ← mkEq e₁ e₂ + if ← isEqv tmp (.const ``False []) then return true + let tmp ← mkHEq e₁ e₂ + isEqv tmp (.const ``False []) + +/-- Is the proposition `e` known to be true? -/ +@[inline] +def isEqTrue (e : Expr) : CCM Bool := + isEqv e (.const ``True []) + +/-- Is the proposition `e` known to be false? -/ +@[inline] +def isEqFalse (e : Expr) : CCM Bool := + isEqv e (.const ``False []) + +/-- Apply transitivity to `H₁` and `H₂`, which are both `Eq` or `HEq` depending on `heqProofs`. -/ +def mkTrans (H₁ H₂ : Expr) (heqProofs : Bool) : MetaM Expr := + if heqProofs then mkHEqTrans H₁ H₂ else mkEqTrans H₁ H₂ + +/-- Apply transitivity to `H₁?` and `H₂`, which are both `Eq` or `HEq` depending on `heqProofs`. + +If `H₁?` is `none`, return `H₂` instead. -/ +def mkTransOpt (H₁? : Option Expr) (H₂ : Expr) (heqProofs : Bool) : MetaM Expr := + match H₁? with + | some H₁ => mkTrans H₁ H₂ heqProofs + | none => pure H₂ + +mutual +/-- Use congruence on arguments to prove `lhs = rhs`. + +That is, tries to prove that `lhsFn lhsArgs[0] ... lhsArgs[n-1] = lhsFn rhsArgs[0] ... rhsArgs[n-1]` +by showing that `lhsArgs[i] = rhsArgs[i]` for all `i`. + +Fails if the head function of `lhs` is not that of `rhs`. -/ +partial def mkCongrProofCore (lhs rhs : Expr) (heqProofs : Bool) : CCM Expr := do + let mut lhsArgsRev : Array Expr := #[] + let mut rhsArgsRev : Array Expr := #[] + let mut lhsIt := lhs + let mut rhsIt := rhs + -- Collect the arguments to `lhs` and `rhs`. + -- As an optimization, we stop collecting arguments as soon as the functions are defeq, + -- so `lhsFn` and `rhsFn` might end up still of the form `(f x y z)` and `(f x' y' z')`. + if lhs != rhs then + repeat + let .app lhsItFn lhsItArg := lhsIt | failure + let .app rhsItFn rhsItArg := rhsIt | failure + lhsArgsRev := lhsArgsRev.push lhsItArg + rhsArgsRev := rhsArgsRev.push rhsItArg + lhsIt := lhsItFn + rhsIt := rhsItFn + if lhsIt == rhsIt then + break + if ← pureIsDefEq lhsIt rhsIt then + break + if ← isEqv lhsIt rhsIt <&&> + inferType lhsIt >>= fun i₁ => inferType rhsIt >>= fun i₂ => pureIsDefEq i₁ i₂ then + break + -- If we collect no arguments, the expressions themselves are defeq; return `rfl`. + if lhsArgsRev.isEmpty then + if heqProofs then + return (← mkHEqRefl lhs) + else + return (← mkEqRefl lhs) + let lhsArgs := lhsArgsRev.reverse + let rhsArgs := rhsArgsRev.reverse + -- Ensure that `lhsFn = rhsFn`, they have the same type and the same list of arguments. + let PLift.up ha ← if ha : lhsArgs.size = rhsArgs.size then pure (PLift.up ha) else failure + let lhsFn := lhsIt + let rhsFn := rhsIt + guard (← isEqv lhsFn rhsFn <||> pureIsDefEq lhsFn rhsFn) + guard (← pureIsDefEq (← inferType lhsFn) (← inferType rhsFn)) + /- Create `r`, a proof for + `lhsFn lhsArgs[0] ... lhsArgs[n-1] = lhsFn rhsArgs[0] ... rhsArgs[n-1]` + where `n := lhsArgs.size` -/ + let some specLemma ← mkCCHCongrTheorem lhsFn lhsArgs.size | failure + let mut kindsIt := specLemma.argKinds + let mut lemmaArgs : Array Expr := #[] + for hi : i in [:lhsArgs.size] do + guard !kindsIt.isEmpty + lemmaArgs := lemmaArgs.push (lhsArgs[i]'hi.2) |>.push (rhsArgs[i]'(ha.symm ▸ hi.2)) + if kindsIt[0]! matches CongrArgKind.heq then + let some p ← getHEqProof (lhsArgs[i]'hi.2) (rhsArgs[i]'(ha.symm ▸ hi.2)) | failure + lemmaArgs := lemmaArgs.push p + else + guard (kindsIt[0]! matches .eq) + let some p ← getEqProof (lhsArgs[i]'hi.2) (rhsArgs[i]'(ha.symm ▸ hi.2)) | failure + lemmaArgs := lemmaArgs.push p + kindsIt := kindsIt.eraseIdx! 0 + let mut r := mkAppN specLemma.proof lemmaArgs + if specLemma.heqResult && !heqProofs then + r ← mkAppM ``eq_of_heq #[r] + else if !specLemma.heqResult && heqProofs then + r ← mkAppM ``heq_of_eq #[r] + if ← pureIsDefEq lhsFn rhsFn then + return r + /- Convert `r` into a proof of `lhs = rhs` using `Eq.rec` and + the proof that `lhsFn = rhsFn` -/ + let some lhsFnEqRhsFn ← getEqProof lhsFn rhsFn | failure + let motive ← + withLocalDeclD `x (← inferType lhsFn) fun x => do + let motiveRhs := mkAppN x rhsArgs + let motive ← if heqProofs then mkHEq lhs motiveRhs else mkEq lhs motiveRhs + let hType ← mkEq lhsFn x + withLocalDeclD `h hType fun h => + mkLambdaFVars #[x, h] motive + mkEqRec motive r lhsFnEqRhsFn + +/-- If `e₁ : R lhs₁ rhs₁`, `e₂ : R lhs₂ rhs₂` and `lhs₁ = rhs₂`, where `R` is a symmetric relation, +prove `R lhs₁ rhs₁` is equivalent to `R lhs₂ rhs₂`. + + * if `lhs₁` is known to equal `lhs₂`, return `none` + * if `lhs₁` is not known to equal `rhs₂`, fail. -/ +partial def mkSymmCongrProof (e₁ e₂ : Expr) (heqProofs : Bool) : CCM (Option Expr) := do + let some (R₁, lhs₁, rhs₁) ← e₁.relSidesIfSymm? | return none + let some (R₂, lhs₂, rhs₂) ← e₂.relSidesIfSymm? | return none + if R₁ != R₂ then return none + if (← isEqv lhs₁ lhs₂) then + return none + guard (← isEqv lhs₁ rhs₂) + /- We must apply symmetry. + The symm congruence table is implicitly using symmetry. + That is, we have + `e₁ := lhs₁ ~R₁~ rhs₁` + and + `e2 := lhs₂ ~R₁~ rhs₂` + But, + `lhs₁ ~R₁~ rhs₂` and `rhs₁ ~R₁~ lhs₂` -/ + /- Given `e₁ := lhs₁ ~R₁~ rhs₁`, + create proof for + `lhs₁ ~R₁~ rhs₁` = `rhs₁ ~R₁~ lhs₁` -/ + let newE₁ ← mkRel R₁ rhs₁ lhs₁ + let e₁IffNewE₁ ← + withLocalDeclD `h₁ e₁ fun h₁ => + withLocalDeclD `h₂ newE₁ fun h₂ => do + mkAppM ``Iff.intro + #[← mkLambdaFVars #[h₁] (← h₁.applySymm), ← mkLambdaFVars #[h₂] (← h₂.applySymm)] + let mut e₁EqNewE₁ := mkApp3 (.const ``propext []) e₁ newE₁ e₁IffNewE₁ + let newE₁EqE₂ ← mkCongrProofCore newE₁ e₂ heqProofs + if heqProofs then + e₁EqNewE₁ ← mkAppM ``heq_of_eq #[e₁EqNewE₁] + return some (← mkTrans e₁EqNewE₁ newE₁EqE₂ heqProofs) + +/-- Use congruence on arguments to prove `e₁ = e₂`. + +Special case: if `e₁` and `e₂` have the form `R lhs₁ rhs₁` and `R lhs₂ rhs₂` such that +`R` is symmetric and `lhs₁ = rhs₂`, then use those facts instead. -/ +partial def mkCongrProof (e₁ e₂ : Expr) (heqProofs : Bool) : CCM Expr := do + if let some r ← mkSymmCongrProof e₁ e₂ heqProofs then + return r + else + mkCongrProofCore e₁ e₂ heqProofs + +/-- Turn a delayed proof into an actual proof term. -/ +partial def mkDelayedProof (H : DelayedExpr) : CCM Expr := do + match H with + | .ofExpr H => return H + | .eqProof lhs rhs => liftOption (← getEqProof lhs rhs) + | .congrArg f h => mkCongrArg f (← mkDelayedProof h) + | .congrFun h a => mkCongrFun (← mkDelayedProof h) (← liftOption a.toExpr) + | .eqSymm h => mkEqSymm (← mkDelayedProof h) + | .eqSymmOpt a₁ a₂ h => + mkAppOptM ``Eq.symm #[none, ← liftOption a₁.toExpr, ← liftOption a₂.toExpr, ← mkDelayedProof h] + | .eqTrans h₁ h₂ => mkEqTrans (← mkDelayedProof h₁) (← mkDelayedProof h₂) + | .eqTransOpt a₁ a₂ a₃ h₁ h₂ => + mkAppOptM ``Eq.trans + #[none, ← liftOption a₁.toExpr, ← liftOption a₂.toExpr, ← liftOption a₃.toExpr, + ← mkDelayedProof h₁, ← mkDelayedProof h₂] + | .heqOfEq h => mkAppM ``heq_of_eq #[← mkDelayedProof h] + | .heqSymm h => mkHEqSymm (← mkDelayedProof h) + +/-- Use the format of `H` to try and construct a proof or `lhs = rhs`: + * If `H = .congr`, then use congruence. + * If `H = .eqTrue`, try to prove `lhs = True` or `rhs = True`, + if they have the format `R a b`, by proving `a = b`. + * Otherwise, return the (delayed) proof encoded by `H` itself. -/ +partial def mkProof (lhs rhs : Expr) (H : EntryExpr) (heqProofs : Bool) : CCM Expr := do + match H with + | .congr => mkCongrProof lhs rhs heqProofs + | .eqTrue => + let (flip, some (R, a, b)) ← + if lhs == .const ``True [] then + ((true, ·)) <$> rhs.relSidesIfRefl? + else + ((false, ·)) <$> lhs.relSidesIfRefl? + | failure + let aRb ← + if R == ``Eq then + getEqProof a b >>= liftOption + else if R == ``HEq then + getHEqProof a b >>= liftOption + else + -- TODO(Leo): the following code assumes R is homogeneous. + -- We should add support arbitrary heterogeneous reflexive relations. + getEqProof a b >>= liftOption >>= fun aEqb => liftM (liftFromEq R aEqb) + let aRbEqTrue ← mkEqTrue aRb + if flip then + mkEqSymm aRbEqTrue + else + return aRbEqTrue + | .refl => + let type ← if heqProofs then mkHEq lhs rhs else mkEq lhs rhs + let proof ← if heqProofs then mkHEqRefl lhs else mkEqRefl lhs + mkExpectedTypeHint proof type + | .ofExpr H => return H + | .ofDExpr H => mkDelayedProof H + +/-- +If `asHEq` is `true`, then build a proof for `HEq e₁ e₂`. +Otherwise, build a proof for `e₁ = e₂`. +The result is `none` if `e₁` and `e₂` are not in the same equivalence class. -/ +partial def getEqProofCore (e₁ e₂ : Expr) (asHEq : Bool) : CCM (Option Expr) := do + if e₁.hasExprMVar || e₂.hasExprMVar then return none + if ← pureIsDefEq e₁ e₂ then + if asHEq then + return some (← mkHEqRefl e₁) + else + return some (← mkEqRefl e₁) + let some n₁ ← getEntry e₁ | return none + let some n₂ ← getEntry e₂ | return none + if n₁.root != n₂.root then return none + let heqProofs ← hasHEqProofs n₁.root + -- 1. Retrieve "path" from `e₁` to `root` + let mut path₁ : Array Expr := #[] + let mut Hs₁ : Array EntryExpr := #[] + let mut visited : RBExprSet := ∅ + let mut it₁ := e₁ + repeat + visited := visited.insert it₁ + let some it₁N ← getEntry it₁ | failure + let some t := it₁N.target | break + path₁ := path₁.push t + let some p := it₁N.proof | failure + Hs₁ := Hs₁.push (← flipProof p it₁N.flipped heqProofs) + it₁ := t + guard (it₁ == n₁.root) + -- 2. The path from `e₂` to root must have at least one element `c` in visited + -- Retrieve "path" from `e₂` to `c` + let mut path₂ : Array Expr := #[] + let mut Hs₂ : Array EntryExpr := #[] + let mut it₂ := e₂ + repeat + if visited.contains it₂ then + break -- found common + let some it₂N ← getEntry it₂ | failure + let some t := it₂N.target | failure + path₂ := path₂.push it₂ + let some p := it₂N.proof | failure + Hs₂ := Hs₂.push (← flipProof p (!it₂N.flipped) heqProofs) + it₂ := t + -- `it₂` is the common element... + -- 3. Shrink `path₁`/`Hs₁` until we find `it₂` (the common element) + repeat + if path₁.isEmpty then + guard (it₂ == e₁) + break + if path₁.back! == it₂ then + -- found it! + break + path₁ := path₁.pop + Hs₁ := Hs₁.pop + + -- 4. Build transitivity proof + let mut pr? : Option Expr := none + let mut lhs := e₁ + for i in [:path₁.size] do + pr? ← some <$> mkTransOpt pr? (← mkProof lhs path₁[i]! Hs₁[i]! heqProofs) heqProofs + lhs := path₁[i]! + let mut i := Hs₂.size + while i > 0 do + i := i - 1 + pr? ← some <$> mkTransOpt pr? (← mkProof lhs path₂[i]! Hs₂[i]! heqProofs) heqProofs + lhs := path₂[i]! + let mut some pr := pr? | failure + if heqProofs && !asHEq then + pr ← mkAppM ``eq_of_heq #[pr] + else if !heqProofs && asHEq then + pr ← mkAppM ``heq_of_eq #[pr] + return pr + +/-- Build a proof for `e₁ = e₂`. +The result is `none` if `e₁` and `e₂` are not in the same equivalence class. -/ +@[inline] +partial def getEqProof (e₁ e₂ : Expr) : CCM (Option Expr) := + getEqProofCore e₁ e₂ false + +/-- Build a proof for `HEq e₁ e₂`. +The result is `none` if `e₁` and `e₂` are not in the same equivalence class. -/ +@[inline] +partial def getHEqProof (e₁ e₂ : Expr) : CCM (Option Expr) := + getEqProofCore e₁ e₂ true +end + +/-- Build a proof for `e = True`. Fails if `e` is not known to be true. -/ +def getEqTrueProof (e : Expr) : CCM Expr := do + guard (← isEqTrue e) + let some p ← getEqProof e (.const ``True []) | failure + return p + +/-- Build a proof for `e = False`. Fails if `e` is not known to be false. -/ +def getEqFalseProof (e : Expr) : CCM Expr := do + guard (← isEqFalse e) + let some p ← getEqProof e (.const ``False []) | failure + return p + +/-- Build a proof for `a = b`. Fails if `a` and `b` are not known to be equal. -/ +def getPropEqProof (a b : Expr) : CCM Expr := do + guard (← isEqv a b) + let some p ← getEqProof a b | failure + return p + +/-- Build a proof of `False` if the context is inconsistent. +Returns `none` if `False` is not known to be true. -/ +def getInconsistencyProof : CCM (Option Expr) := do + guard !(← get).frozePartitions + if let some p ← getEqProof (.const ``True []) (.const ``False []) then + return some (← mkAppM ``false_of_true_eq_false #[p]) + else + return none + +/-- Given `a`, `a₁` and `a₁NeB : a₁ ≠ b`, return a proof of `a ≠ b` if `a` and `a₁` are in the +same equivalence class. -/ +def mkNeOfEqOfNe (a a₁ a₁NeB : Expr) : CCM (Option Expr) := do + guard (← isEqv a a₁) + if a == a₁ then + return some a₁NeB + let aEqA₁ ← getEqProof a a₁ + match aEqA₁ with + | none => return none -- failed to build proof + | some aEqA₁ => mkAppM ``ne_of_eq_of_ne #[aEqA₁, a₁NeB] + +/-- Given `aNeB₁ : a ≠ b₁`, `b₁` and `b`, return a proof of `a ≠ b` if `b` and `b₁` are in the +same equivalence class. -/ +def mkNeOfNeOfEq (aNeB₁ b₁ b : Expr) : CCM (Option Expr) := do + guard (← isEqv b b₁) + if b == b₁ then + return some aNeB₁ + let b₁EqB ← getEqProof b b₁ + match b₁EqB with + | none => return none -- failed to build proof + | some b₁EqB => mkAppM ``ne_of_ne_of_eq #[aNeB₁, b₁EqB] + +/-- Return the proof of `e₁ = e₂` using `ac_rfl` tactic. -/ +def mkACProof (e₁ e₂ : Expr) : MetaM Expr := do + let eq ← mkEq e₁ e₂ + let .mvar m ← mkFreshExprSyntheticOpaqueMVar eq | failure + AC.rewriteUnnormalizedRefl m + let pr ← instantiateMVars (.mvar m) + mkExpectedTypeHint pr eq + +/-- Given `tr := t*r` `sr := s*r` `tEqs : t = s`, return a proof for `tr = sr` + + We use `a*b` to denote an AC application. That is, `(a*b)*(c*a)` is the term `a*a*b*c`. -/ +def mkACSimpProof (tr t s r sr : ACApps) (tEqs : DelayedExpr) : MetaM DelayedExpr := do + if tr == t then + return tEqs + else if tr == sr then + let some tre := tr.toExpr | failure + DelayedExpr.ofExpr <$> mkEqRefl tre + else + let .apps op _ := tr | failure + let some re := r.toExpr | failure + let some te := t.toExpr | failure + let some se := s.toExpr | failure + let some tre := tr.toExpr | failure + let some sre := sr.toExpr | failure + let opr := op.app re -- `(*) r` + let rt := mkApp2 op re te -- `r * t` + let rs := mkApp2 op re se -- `r * s` + let rtEqrs := DelayedExpr.congrArg opr tEqs + let trEqrt ← mkACProof tre rt + let rsEqsr ← mkACProof rs sre + return .eqTrans (.eqTrans trEqrt rtEqrs) rsEqsr + +/-- Given `e := lhs * r` and `H : lhs = rhs`, return `rhs * r` and the proof of `e = rhs * r`. -/ +def simplifyACCore (e lhs rhs : ACApps) (H : DelayedExpr) : + CCM (ACApps × DelayedExpr) := do + guard (lhs.isSubset e) + if e == lhs then + return (rhs, H) + else + let .apps op _ := e | failure + let newArgs := e.diff lhs + let r : ACApps := if newArgs.isEmpty then default else .mkApps op newArgs + let newArgs := ACApps.append op rhs newArgs + let newE := ACApps.mkApps op newArgs + let some true := (← get).opInfo.find? op | failure + let newPr ← mkACSimpProof e lhs rhs r newE H + return (newE, newPr) + +/-- The single step of `simplifyAC`. + +Simplifies an expression `e` by either simplifying one argument to the AC operator, or the whole +expression. -/ +def simplifyACStep (e : ACApps) : CCM (Option (ACApps × DelayedExpr)) := do + if let .apps _ args := e then + for h : i in [:args.size] do + if i == 0 || (args[i]'h.2) != (args[i - 1]'(Nat.lt_of_le_of_lt (i.sub_le 1) h.2)) then + let some ae := (← get).acEntries.find? (args[i]'h.2) | failure + let occs := ae.RLHSOccs + let mut Rlhs? : Option ACApps := none + for Rlhs in occs do + if Rlhs.isSubset e then + Rlhs? := some Rlhs + break + if let some Rlhs := Rlhs? then + let some (Rrhs, H) := (← get).acR.find? Rlhs | failure + return (some <| ← simplifyACCore e Rlhs Rrhs H) + else if let some p := (← get).acR.find? e then + return some p + return none + +/-- If `e` can be simplified by the AC module, return the simplified term and the proof term of the +equality. -/ +def simplifyAC (e : ACApps) : CCM (Option (ACApps × DelayedExpr)) := do + let mut some (curr, pr) ← simplifyACStep e | return none + repeat + let some (newCurr, newPr) ← simplifyACStep curr | break + pr := .eqTransOpt e curr newCurr pr newPr + curr := newCurr + return some (curr, pr) + +end Mathlib.Tactic.CC.CCM diff --git a/scripts/noshake.json b/scripts/noshake.json index acb992ed41e28..3ae49049109e5 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -300,7 +300,7 @@ "Mathlib.Tactic.CategoryTheory.BicategoryCoherence": ["Mathlib.CategoryTheory.Bicategory.Coherence", "Mathlib.Tactic.CategoryTheory.BicategoricalComp"], - "Mathlib.Tactic.CC.Addition": + "Mathlib.Tactic.CC.MkProof": ["Mathlib.Logic.Basic", "Mathlib.Tactic.CC.Lemmas"], "Mathlib.Tactic.ByContra": ["Batteries.Tactic.Init"], "Mathlib.Tactic.Bound.Init": ["Aesop.Frontend.Command"], From 79dae4b2147c58f873a92001dcfd20688afd8618 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 18 Dec 2024 13:02:53 +0000 Subject: [PATCH 765/829] chore(Analysis): rename `closedUnitBall` and `closed_unit_ball` to `unitClosedBall` (#9755) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Since `closedBall` is a function name, it should appear as an unbroken token in the lemma name [Zulip poll](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/unit_closedBall.20or.20unitClosedBall/near/412926794) Co-authored-by: Yaël Dillies --- Mathlib/Analysis/BoxIntegral/Basic.lean | 2 +- Mathlib/Analysis/CStarAlgebra/Multiplier.lean | 8 ++++---- .../Analysis/CStarAlgebra/Unitization.lean | 8 ++++---- .../Analysis/Convex/StrictConvexSpace.lean | 20 +++++++++++-------- Mathlib/Analysis/Normed/Field/Basic.lean | 5 ++++- Mathlib/Analysis/Normed/Field/Lemmas.lean | 2 +- .../NormedSpace/OperatorNorm/NNNorm.lean | 12 ++++++++--- Mathlib/Analysis/NormedSpace/Pointwise.lean | 13 ++++++++---- Mathlib/Analysis/NormedSpace/Real.lean | 5 ++++- .../Measure/Lebesgue/EqHaar.lean | 7 +++++-- 10 files changed, 53 insertions(+), 29 deletions(-) diff --git a/Mathlib/Analysis/BoxIntegral/Basic.lean b/Mathlib/Analysis/BoxIntegral/Basic.lean index 44c6028cdd79d..b4a22b6102a0b 100644 --- a/Mathlib/Analysis/BoxIntegral/Basic.lean +++ b/Mathlib/Analysis/BoxIntegral/Basic.lean @@ -745,7 +745,7 @@ theorem integrable_of_continuousOn [CompleteSpace E] {I : Box ι} {f : ℝⁿ · obtain ⟨C, hC⟩ := (NormedSpace.isBounded_iff_subset_smul_closedBall ℝ).1 (I.isCompact_Icc.image_of_continuousOn hc).isBounded use ‖C‖, fun x hx ↦ by - simpa only [smul_closedUnitBall, mem_closedBall_zero_iff] using hC (Set.mem_image_of_mem f hx) + simpa only [smul_unitClosedBall, mem_closedBall_zero_iff] using hC (Set.mem_image_of_mem f hx) · refine eventually_of_mem ?_ (fun x hx ↦ hc.continuousWithinAt hx) rw [mem_ae_iff, μ.restrict_apply] <;> simp [MeasurableSet.compl_iff.2 I.measurableSet_Icc] diff --git a/Mathlib/Analysis/CStarAlgebra/Multiplier.lean b/Mathlib/Analysis/CStarAlgebra/Multiplier.lean index 8f62eac684540..900b2eb10e19f 100644 --- a/Mathlib/Analysis/CStarAlgebra/Multiplier.lean +++ b/Mathlib/Analysis/CStarAlgebra/Multiplier.lean @@ -630,7 +630,7 @@ instance instCStarRing : CStarRing 𝓜(𝕜, A) where On the other hand, for any `‖z‖ ≤ 1`, we may choose `x := star z` and `y := z` to get: `‖star (L (star x)) * L y‖ = ‖star (L z) * (L z)‖ = ‖L z‖ ^ 2`, and taking the supremum over all such `z` yields that the supremum is at least `‖L‖ ^ 2`. It is the latter part of the - argument where `DenselyNormedField 𝕜` is required (for `sSup_closed_unit_ball_eq_nnnorm`). -/ + argument where `DenselyNormedField 𝕜` is required (for `sSup_unitClosedBall_eq_nnnorm`). -/ have hball : (Metric.closedBall (0 : A) 1).Nonempty := Metric.nonempty_closedBall.2 zero_le_one have key : @@ -645,9 +645,9 @@ instance instCStarRing : CStarRing 𝓜(𝕜, A) where (a.fst.le_opNorm_of_le hy)) _ ≤ ‖a‖₊ * ‖a‖₊ := by simp only [mul_one, nnnorm_fst, le_rfl] rw [← nnnorm_snd] - simp only [mul_snd, ← sSup_closed_unit_ball_eq_nnnorm, star_snd, mul_apply] + simp only [mul_snd, ← sSup_unitClosedBall_eq_nnnorm, star_snd, mul_apply] simp only [← @opNNNorm_mul_apply 𝕜 _ A] - simp only [← sSup_closed_unit_ball_eq_nnnorm, mul_apply'] + simp only [← sSup_unitClosedBall_eq_nnnorm, mul_apply'] refine csSup_eq_of_forall_le_of_forall_lt_exists_gt (hball.image _) ?_ fun r hr => ?_ · rintro - ⟨x, hx, rfl⟩ refine csSup_le (hball.image _) ?_ @@ -655,7 +655,7 @@ instance instCStarRing : CStarRing 𝓜(𝕜, A) where exact key x y (mem_closedBall_zero_iff.1 hx) (mem_closedBall_zero_iff.1 hy) · simp only [Set.mem_image, Set.mem_setOf_eq, exists_prop, exists_exists_and_eq_and] have hr' : NNReal.sqrt r < ‖a‖₊ := ‖a‖₊.sqrt_mul_self ▸ NNReal.sqrt_lt_sqrt.2 hr - simp_rw [← nnnorm_fst, ← sSup_closed_unit_ball_eq_nnnorm] at hr' + simp_rw [← nnnorm_fst, ← sSup_unitClosedBall_eq_nnnorm] at hr' obtain ⟨_, ⟨x, hx, rfl⟩, hxr⟩ := exists_lt_of_lt_csSup (hball.image _) hr' have hx' : ‖x‖₊ ≤ 1 := mem_closedBall_zero_iff.1 hx refine ⟨star x, mem_closedBall_zero_iff.2 ((nnnorm_star x).trans_le hx'), ?_⟩ diff --git a/Mathlib/Analysis/CStarAlgebra/Unitization.lean b/Mathlib/Analysis/CStarAlgebra/Unitization.lean index 94c73b1a5216b..8771d6ba22cae 100644 --- a/Mathlib/Analysis/CStarAlgebra/Unitization.lean +++ b/Mathlib/Analysis/CStarAlgebra/Unitization.lean @@ -60,7 +60,7 @@ variable (E) instance CStarRing.instRegularNormedAlgebra : RegularNormedAlgebra 𝕜 E where isometry_mul' := AddMonoidHomClass.isometry_of_norm (mul 𝕜 E) fun a => NNReal.eq_iff.mp <| show ‖mul 𝕜 E a‖₊ = ‖a‖₊ by - rw [← sSup_closed_unit_ball_eq_nnnorm] + rw [← sSup_unitClosedBall_eq_nnnorm] refine csSup_eq_of_forall_le_of_forall_lt_exists_gt ?_ ?_ fun r hr => ?_ · exact (Metric.nonempty_closedBall.mpr zero_le_one).image _ · rintro - ⟨x, hx, rfl⟩ @@ -87,12 +87,12 @@ variable {E} out so that declaring the `CStarRing` instance doesn't time out. -/ theorem Unitization.norm_splitMul_snd_sq (x : Unitization 𝕜 E) : ‖(Unitization.splitMul 𝕜 E x).snd‖ ^ 2 ≤ ‖(Unitization.splitMul 𝕜 E (star x * x)).snd‖ := by - /- The key idea is that we can use `sSup_closed_unit_ball_eq_norm` to make this about + /- The key idea is that we can use `sSup_unitClosedBall_eq_norm` to make this about applying this linear map to elements of norm at most one. There is a bit of `sqrt` and `sq` shuffling that needs to occur, which is primarily just an annoyance. -/ refine (Real.le_sqrt (norm_nonneg _) (norm_nonneg _)).mp ?_ simp only [Unitization.splitMul_apply] - rw [← sSup_closed_unit_ball_eq_norm] + rw [← sSup_unitClosedBall_eq_norm] refine csSup_le ((Metric.nonempty_closedBall.2 zero_le_one).image _) ?_ rintro - ⟨b, hb, rfl⟩ simp only @@ -101,7 +101,7 @@ theorem Unitization.norm_splitMul_snd_sq (x : Unitization 𝕜 E) : ← CStarRing.norm_star_mul_self, ContinuousLinearMap.add_apply, star_add, mul_apply', Algebra.algebraMap_eq_smul_one, ContinuousLinearMap.smul_apply, ContinuousLinearMap.one_apply, star_mul, star_smul, add_mul, smul_mul_assoc, ← mul_smul_comm, - mul_assoc, ← mul_add, ← sSup_closed_unit_ball_eq_norm] + mul_assoc, ← mul_add, ← sSup_unitClosedBall_eq_norm] refine (norm_mul_le _ _).trans ?_ calc _ ≤ ‖star x.fst • (x.fst • b + x.snd * b) + star x.snd * (x.fst • b + x.snd * b)‖ := by diff --git a/Mathlib/Analysis/Convex/StrictConvexSpace.lean b/Mathlib/Analysis/Convex/StrictConvexSpace.lean index 32a1cef9cd3a9..9e50330e7e3d1 100644 --- a/Mathlib/Analysis/Convex/StrictConvexSpace.lean +++ b/Mathlib/Analysis/Convex/StrictConvexSpace.lean @@ -35,7 +35,7 @@ In a strictly convex space, we prove We also provide several lemmas that can be used as alternative constructors for `StrictConvex ℝ E`: -- `StrictConvexSpace.of_strictConvex_closed_unit_ball`: if `closed_ball (0 : E) 1` is strictly +- `StrictConvexSpace.of_strictConvex_unitClosedBall`: if `closed_ball (0 : E) 1` is strictly convex, then `E` is a strictly convex space; - `StrictConvexSpace.of_norm_add`: if `‖x + y‖ = ‖x‖ + ‖y‖` implies `SameRay ℝ x y` for all @@ -57,7 +57,7 @@ open Convex Pointwise Set Metric require balls of positive radius with center at the origin to be strictly convex in the definition, then prove that any closed ball is strictly convex in `strictConvex_closedBall` below. -See also `StrictConvexSpace.of_strictConvex_closed_unit_ball`. -/ +See also `StrictConvexSpace.of_strictConvex_unitClosedBall`. -/ class StrictConvexSpace (𝕜 E : Type*) [NormedLinearOrderedField 𝕜] [NormedAddCommGroup E] [NormedSpace 𝕜 E] : Prop where strictConvex_closedBall : ∀ r : ℝ, 0 < r → StrictConvex 𝕜 (closedBall (0 : E) r) @@ -76,9 +76,13 @@ theorem strictConvex_closedBall [StrictConvexSpace 𝕜 E] (x : E) (r : ℝ) : variable [NormedSpace ℝ E] /-- A real normed vector space is strictly convex provided that the unit ball is strictly convex. -/ -theorem StrictConvexSpace.of_strictConvex_closed_unit_ball [LinearMap.CompatibleSMul E E 𝕜 ℝ] +theorem StrictConvexSpace.of_strictConvex_unitClosedBall [LinearMap.CompatibleSMul E E 𝕜 ℝ] (h : StrictConvex 𝕜 (closedBall (0 : E) 1)) : StrictConvexSpace 𝕜 E := - ⟨fun r hr => by simpa only [smul_closedUnitBall_of_nonneg hr.le] using h.smul r⟩ + ⟨fun r hr => by simpa only [smul_unitClosedBall_of_nonneg hr.le] using h.smul r⟩ + +@[deprecated (since := "2024-12-01")] +alias StrictConvexSpace.of_strictConvex_closed_unit_ball := + StrictConvexSpace.of_strictConvex_unitClosedBall /-- Strict convexity is equivalent to `‖a • x + b • y‖ < 1` for all `x` and `y` of norm at most `1` and all strictly positive `a` and `b` such that `a + b = 1`. This lemma shows that it suffices to @@ -87,7 +91,7 @@ theorem StrictConvexSpace.of_norm_combo_lt_one (h : ∀ x y : E, ‖x‖ = 1 → ‖y‖ = 1 → x ≠ y → ∃ a b : ℝ, a + b = 1 ∧ ‖a • x + b • y‖ < 1) : StrictConvexSpace ℝ E := by refine - StrictConvexSpace.of_strictConvex_closed_unit_ball ℝ + StrictConvexSpace.of_strictConvex_unitClosedBall ℝ ((convex_closedBall _ _).strictConvex' fun x hx y hy hne => ?_) rw [interior_closedBall (0 : E) one_ne_zero, closedBall_diff_ball, mem_sphere_zero_iff_norm] at hx hy @@ -101,7 +105,7 @@ theorem StrictConvexSpace.of_norm_combo_ne_one ∀ x y : E, ‖x‖ = 1 → ‖y‖ = 1 → x ≠ y → ∃ a b : ℝ, 0 ≤ a ∧ 0 ≤ b ∧ a + b = 1 ∧ ‖a • x + b • y‖ ≠ 1) : StrictConvexSpace ℝ E := by - refine StrictConvexSpace.of_strictConvex_closed_unit_ball ℝ + refine StrictConvexSpace.of_strictConvex_unitClosedBall ℝ ((convex_closedBall _ _).strictConvex ?_) simp only [interior_closedBall _ one_ne_zero, closedBall_diff_ball, Set.Pairwise, frontier_closedBall _ one_ne_zero, mem_sphere_zero_iff_norm] @@ -164,8 +168,8 @@ theorem norm_add_lt_of_not_sameRay (h : ¬SameRay ℝ x y) : ‖x + y‖ < ‖x rw [← norm_pos_iff] at hx hy have hxy : 0 < ‖x‖ + ‖y‖ := add_pos hx hy have := - combo_mem_ball_of_ne (inv_norm_smul_mem_closed_unit_ball x) - (inv_norm_smul_mem_closed_unit_ball y) hne (div_pos hx hxy) (div_pos hy hxy) + combo_mem_ball_of_ne (inv_norm_smul_mem_unitClosedBall x) + (inv_norm_smul_mem_unitClosedBall y) hne (div_pos hx hxy) (div_pos hy hxy) (by rw [← add_div, div_self hxy.ne']) rwa [mem_ball_zero_iff, div_eq_inv_mul, div_eq_inv_mul, mul_smul, mul_smul, smul_inv_smul₀ hx.ne', smul_inv_smul₀ hy.ne', ← smul_add, norm_smul, Real.norm_of_nonneg (inv_pos.2 hxy).le, ← diff --git a/Mathlib/Analysis/Normed/Field/Basic.lean b/Mathlib/Analysis/Normed/Field/Basic.lean index d730d9d8baf82..bd7716e46f8c9 100644 --- a/Mathlib/Analysis/Normed/Field/Basic.lean +++ b/Mathlib/Analysis/Normed/Field/Basic.lean @@ -759,10 +759,13 @@ lemma norm_le_one_of_discrete · simp · simp [norm_eq_one_iff_ne_zero_of_discrete.mpr hx] -lemma discreteTopology_unit_closedBall_eq_univ : (Metric.closedBall 0 1 : Set 𝕜) = Set.univ := by +lemma unitClosedBall_eq_univ_of_discrete : (Metric.closedBall 0 1 : Set 𝕜) = Set.univ := by ext simp +@[deprecated (since := "2024-12-01")] +alias discreteTopology_unit_closedBall_eq_univ := unitClosedBall_eq_univ_of_discrete + end Discrete end NormedDivisionRing diff --git a/Mathlib/Analysis/Normed/Field/Lemmas.lean b/Mathlib/Analysis/Normed/Field/Lemmas.lean index 2943b39e7e097..37916052cd845 100644 --- a/Mathlib/Analysis/Normed/Field/Lemmas.lean +++ b/Mathlib/Analysis/Normed/Field/Lemmas.lean @@ -429,7 +429,7 @@ lemma NormedField.completeSpace_iff_isComplete_closedBall {K : Type*} [NormedFie · exact Metric.isClosed_ball.isComplete rcases NormedField.discreteTopology_or_nontriviallyNormedField K with _|⟨_, rfl⟩ · rwa [completeSpace_iff_isComplete_univ, - ← NormedDivisionRing.discreteTopology_unit_closedBall_eq_univ] + ← NormedDivisionRing.unitClosedBall_eq_univ_of_discrete] refine Metric.complete_of_cauchySeq_tendsto fun u hu ↦ ?_ obtain ⟨k, hk⟩ := hu.norm_bddAbove have kpos : 0 ≤ k := (_root_.norm_nonneg (u 0)).trans (hk (by simp)) diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/NNNorm.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/NNNorm.lean index 020de596c8f44..51f3b1705bfc9 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/NNNorm.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/NNNorm.lean @@ -190,7 +190,7 @@ theorem sSup_unit_ball_eq_norm {𝕜 𝕜₂ E F : Type*} [NormedAddCommGroup E] sSup ((fun x => ‖f x‖) '' ball 0 1) = ‖f‖ := by simpa only [NNReal.coe_sSup, Set.image_image] using NNReal.coe_inj.2 f.sSup_unit_ball_eq_nnnorm -theorem sSup_closed_unit_ball_eq_nnnorm {𝕜 𝕜₂ E F : Type*} [NormedAddCommGroup E] +theorem sSup_unitClosedBall_eq_nnnorm {𝕜 𝕜₂ E F : Type*} [NormedAddCommGroup E] [SeminormedAddCommGroup F] [DenselyNormedField 𝕜] [NontriviallyNormedField 𝕜₂] {σ₁₂ : 𝕜 →+* 𝕜₂} [NormedSpace 𝕜 E] [NormedSpace 𝕜₂ F] [RingHomIsometric σ₁₂] (f : E →SL[σ₁₂] F) : sSup ((fun x => ‖f x‖₊) '' closedBall 0 1) = ‖f‖₊ := by @@ -202,12 +202,18 @@ theorem sSup_closed_unit_ball_eq_nnnorm {𝕜 𝕜₂ E F : Type*} [NormedAddCom exact csSup_le_csSup ⟨‖f‖₊, hbdd⟩ ((nonempty_ball.2 zero_lt_one).image _) (Set.image_subset _ ball_subset_closedBall) -theorem sSup_closed_unit_ball_eq_norm {𝕜 𝕜₂ E F : Type*} [NormedAddCommGroup E] +@[deprecated (since := "2024-12-01")] +alias sSup_closed_unit_ball_eq_nnnorm := sSup_unitClosedBall_eq_nnnorm + +theorem sSup_unitClosedBall_eq_norm {𝕜 𝕜₂ E F : Type*} [NormedAddCommGroup E] [SeminormedAddCommGroup F] [DenselyNormedField 𝕜] [NontriviallyNormedField 𝕜₂] {σ₁₂ : 𝕜 →+* 𝕜₂} [NormedSpace 𝕜 E] [NormedSpace 𝕜₂ F] [RingHomIsometric σ₁₂] (f : E →SL[σ₁₂] F) : sSup ((fun x => ‖f x‖) '' closedBall 0 1) = ‖f‖ := by simpa only [NNReal.coe_sSup, Set.image_image] using - NNReal.coe_inj.2 f.sSup_closed_unit_ball_eq_nnnorm + NNReal.coe_inj.2 f.sSup_unitClosedBall_eq_nnnorm + +@[deprecated (since := "2024-12-01")] +alias sSup_closed_unit_ball_eq_norm := sSup_unitClosedBall_eq_norm end Sup diff --git a/Mathlib/Analysis/NormedSpace/Pointwise.lean b/Mathlib/Analysis/NormedSpace/Pointwise.lean index 866ca33561073..fb81f7ab0111d 100644 --- a/Mathlib/Analysis/NormedSpace/Pointwise.lean +++ b/Mathlib/Analysis/NormedSpace/Pointwise.lean @@ -361,16 +361,21 @@ theorem smul_closedBall (c : 𝕜) (x : E) {r : ℝ} (hr : 0 ≤ r) : · simp [hr, zero_smul_set, Set.singleton_zero, nonempty_closedBall] · exact smul_closedBall' hc x r -theorem smul_closedUnitBall (c : 𝕜) : c • closedBall (0 : E) (1 : ℝ) = closedBall (0 : E) ‖c‖ := by +theorem smul_unitClosedBall (c : 𝕜) : c • closedBall (0 : E) (1 : ℝ) = closedBall (0 : E) ‖c‖ := by rw [_root_.smul_closedBall _ _ zero_le_one, smul_zero, mul_one] +@[deprecated (since := "2024-12-01")] alias smul_closedUnitBall := smul_unitClosedBall + variable [NormedSpace ℝ E] /-- In a real normed space, the image of the unit closed ball under multiplication by a nonnegative number `r` is the closed ball of radius `r` with center at the origin. -/ -theorem smul_closedUnitBall_of_nonneg {r : ℝ} (hr : 0 ≤ r) : +theorem smul_unitClosedBall_of_nonneg {r : ℝ} (hr : 0 ≤ r) : r • closedBall (0 : E) 1 = closedBall (0 : E) r := by - rw [smul_closedUnitBall, Real.norm_of_nonneg hr] + rw [smul_unitClosedBall, Real.norm_of_nonneg hr] + +@[deprecated (since := "2024-12-01")] +alias smul_closedUnitBall_of_nonneg := smul_unitClosedBall_of_nonneg /-- In a nontrivial real normed space, a sphere is nonempty if and only if its radius is nonnegative. -/ @@ -400,6 +405,6 @@ theorem affinity_unitBall {r : ℝ} (hr : 0 < r) (x : E) : x +ᵥ r • ball (0 `fun y ↦ x + r • y`. -/ theorem affinity_unitClosedBall {r : ℝ} (hr : 0 ≤ r) (x : E) : x +ᵥ r • closedBall (0 : E) 1 = closedBall x r := by - rw [smul_closedUnitBall, Real.norm_of_nonneg hr, vadd_closedBall_zero] + rw [smul_unitClosedBall, Real.norm_of_nonneg hr, vadd_closedBall_zero] end NormedAddCommGroup diff --git a/Mathlib/Analysis/NormedSpace/Real.lean b/Mathlib/Analysis/NormedSpace/Real.lean index 09c0aa69993d6..aed2117e262d0 100644 --- a/Mathlib/Analysis/NormedSpace/Real.lean +++ b/Mathlib/Analysis/NormedSpace/Real.lean @@ -35,11 +35,14 @@ section Seminormed variable {E : Type*} [SeminormedAddCommGroup E] [NormedSpace ℝ E] -theorem inv_norm_smul_mem_closed_unit_ball (x : E) : +theorem inv_norm_smul_mem_unitClosedBall (x : E) : ‖x‖⁻¹ • x ∈ closedBall (0 : E) 1 := by simp only [mem_closedBall_zero_iff, norm_smul, norm_inv, norm_norm, ← div_eq_inv_mul, div_self_le_one] +@[deprecated (since := "2024-12-01")] +alias inv_norm_smul_mem_closed_unit_ball := inv_norm_smul_mem_unitClosedBall + theorem norm_smul_of_nonneg {t : ℝ} (ht : 0 ≤ t) (x : E) : ‖t • x‖ = t * ‖x‖ := by rw [norm_smul, Real.norm_eq_abs, abs_of_nonneg ht] diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean index 8a0f8da0ef7ad..d0e675c84de16 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean @@ -461,7 +461,7 @@ theorem addHaar_closedBall' (x : E) {r : ℝ} (hr : 0 ≤ r) : μ (closedBall x r) = ENNReal.ofReal (r ^ finrank ℝ E) * μ (closedBall 0 1) := by rw [← addHaar_closedBall_mul μ x hr zero_le_one, mul_one] -theorem addHaar_closed_unit_ball_eq_addHaar_unit_ball : +theorem addHaar_unitClosedBall_eq_addHaar_unitBall : μ (closedBall (0 : E) 1) = μ (ball 0 1) := by apply le_antisymm _ (measure_mono ball_subset_closedBall) have A : Tendsto @@ -476,9 +476,12 @@ theorem addHaar_closed_unit_ball_eq_addHaar_unit_ball : rw [← addHaar_closedBall' μ (0 : E) hr.1.le] exact measure_mono (closedBall_subset_ball hr.2) +@[deprecated (since := "2024-12-01")] +alias addHaar_closed_unit_ball_eq_addHaar_unit_ball := addHaar_unitClosedBall_eq_addHaar_unitBall + theorem addHaar_closedBall (x : E) {r : ℝ} (hr : 0 ≤ r) : μ (closedBall x r) = ENNReal.ofReal (r ^ finrank ℝ E) * μ (ball 0 1) := by - rw [addHaar_closedBall' μ x hr, addHaar_closed_unit_ball_eq_addHaar_unit_ball] + rw [addHaar_closedBall' μ x hr, addHaar_unitClosedBall_eq_addHaar_unitBall] theorem addHaar_closedBall_eq_addHaar_ball [Nontrivial E] (x : E) (r : ℝ) : μ (closedBall x r) = μ (ball x r) := by From cbe8d43a7e6c614db88fc2ad8f0520e50c1d3dae Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Wed, 18 Dec 2024 13:24:51 +0000 Subject: [PATCH 766/829] feat: add some iteratedDeriv variants of iteratedDerivWithin lemmas (#19849) Based on work from EulerMaclaurin. Co-authored-by: Geoffrey Irving --- .../Calculus/IteratedDeriv/Lemmas.lean | 44 ++++++++++++++++--- 1 file changed, 39 insertions(+), 5 deletions(-) diff --git a/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean b/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean index 0a4e4120acb15..1c58f1a462e70 100644 --- a/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean +++ b/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean @@ -52,7 +52,7 @@ theorem iteratedDerivWithin_const_add (hn : 0 < n) (c : F) : intro y hy exact derivWithin_const_add (h.uniqueDiffWithinAt hy) _ -theorem iteratedDerivWithin_const_neg (hn : 0 < n) (c : F) : +theorem iteratedDerivWithin_const_sub (hn : 0 < n) (c : F) : iteratedDerivWithin n (fun z => c - f z) s x = iteratedDerivWithin n (fun z => -f z) s x := by obtain ⟨n, rfl⟩ := n.exists_eq_succ_of_ne_zero hn.ne' rw [iteratedDerivWithin_succ' h hx, iteratedDerivWithin_succ' h hx] @@ -62,12 +62,19 @@ theorem iteratedDerivWithin_const_neg (hn : 0 < n) (c : F) : rw [derivWithin.neg this] exact derivWithin_const_sub this _ +@[deprecated (since := "2024-12-10")] +alias iteratedDerivWithin_const_neg := iteratedDerivWithin_const_sub + +/-- Note: this is unrelated to `iteratedDeriv_const_smul`, where the scalar multiplication works on +the domain. -/ theorem iteratedDerivWithin_const_smul (c : R) (hf : ContDiffOn 𝕜 n f s) : iteratedDerivWithin n (c • f) s x = c • iteratedDerivWithin n f s x := by simp_rw [iteratedDerivWithin] rw [iteratedFDerivWithin_const_smul_apply hf h hx] simp only [ContinuousMultilinearMap.smul_apply] +/-- Note: this is unrelated to `iteratedDeriv_const_mul`, where the multiplication works on the +domain. -/ theorem iteratedDerivWithin_const_mul (c : 𝕜) {f : 𝕜 → 𝕜} (hf : ContDiffOn 𝕜 n f s) : iteratedDerivWithin n (fun z => c * f z) s x = c * iteratedDerivWithin n f s x := by simpa using iteratedDerivWithin_const_smul (F := 𝕜) hx h c hf @@ -91,6 +98,35 @@ theorem iteratedDerivWithin_sub (hf : ContDiffOn 𝕜 n f s) (hg : ContDiffOn end +lemma iteratedDeriv_add (hf : ContDiff 𝕜 n f) (hg : ContDiff 𝕜 n g) : + iteratedDeriv n (f + g) x = iteratedDeriv n f x + iteratedDeriv n g x := by + simpa only [iteratedDerivWithin_univ] using + iteratedDerivWithin_add (Set.mem_univ _) uniqueDiffOn_univ + (contDiffOn_univ.mpr hf) (contDiffOn_univ.mpr hg) + +theorem iteratedDeriv_const_add (hn : 0 < n) (c : F) : + iteratedDeriv n (fun z => c + f z) x = iteratedDeriv n f x := by + simpa only [iteratedDerivWithin_univ] using + iteratedDerivWithin_const_add (Set.mem_univ _) uniqueDiffOn_univ hn c + +theorem iteratedDeriv_const_sub (hn : 0 < n) (c : F) : + iteratedDeriv n (fun z => c - f z) x = iteratedDeriv n (-f) x := by + simpa only [iteratedDerivWithin_univ] using + iteratedDerivWithin_const_sub (Set.mem_univ _) uniqueDiffOn_univ hn c + +lemma iteratedDeriv_neg (n : ℕ) (f : 𝕜 → F) (a : 𝕜) : + iteratedDeriv n (fun x ↦ -(f x)) a = -(iteratedDeriv n f a) := by + simpa only [iteratedDerivWithin_univ] using + iteratedDerivWithin_neg (Set.mem_univ a) uniqueDiffOn_univ f + +lemma iteratedDeriv_sub (hf : ContDiff 𝕜 n f) (hg : ContDiff 𝕜 n g) : + iteratedDeriv n (f - g) x = iteratedDeriv n f x - iteratedDeriv n g x := by + simpa only [iteratedDerivWithin_univ] using + iteratedDerivWithin_sub (Set.mem_univ _) uniqueDiffOn_univ + (contDiffOn_univ.mpr hf) (contDiffOn_univ.mpr hg) + +/-- Note: this is unrelated to `iteratedDerivWithin_const_smul`, where the scalar multiplication +works on the codomain. -/ theorem iteratedDeriv_const_smul {n : ℕ} {f : 𝕜 → F} (h : ContDiff 𝕜 n f) (c : 𝕜) : iteratedDeriv n (fun x => f (c * x)) = fun x => c ^ n • iteratedDeriv n f (c * x) := by induction n with @@ -108,14 +144,12 @@ theorem iteratedDeriv_const_smul {n : ℕ} {f : 𝕜 → F} (h : ContDiff 𝕜 n ← Function.comp_def, deriv.scomp x h₀ (differentiableAt_id'.const_mul _), deriv_const_mul _ differentiableAt_id', deriv_id'', smul_smul, mul_one, pow_succ] +/-- Note: this is unrelated to `iteratedDerivWithin_const_mul`, where the multiplication works on +the codomain. -/ theorem iteratedDeriv_const_mul {n : ℕ} {f : 𝕜 → 𝕜} (h : ContDiff 𝕜 n f) (c : 𝕜) : iteratedDeriv n (fun x => f (c * x)) = fun x => c ^ n * iteratedDeriv n f (c * x) := by simpa only [smul_eq_mul] using iteratedDeriv_const_smul h c -lemma iteratedDeriv_neg (n : ℕ) (f : 𝕜 → F) (a : 𝕜) : - iteratedDeriv n (fun x ↦ -(f x)) a = -(iteratedDeriv n f a) := by - simp_rw [← iteratedDerivWithin_univ, iteratedDerivWithin_neg' (Set.mem_univ a) uniqueDiffOn_univ] - lemma iteratedDeriv_comp_neg (n : ℕ) (f : 𝕜 → F) (a : 𝕜) : iteratedDeriv n (fun x ↦ f (-x)) a = (-1 : 𝕜) ^ n • iteratedDeriv n f (-a) := by induction' n with n ih generalizing a From 211dc533566a173988a01067bb50d7917a097103 Mon Sep 17 00:00:00 2001 From: Xavier Roblot Date: Wed, 18 Dec 2024 15:53:51 +0000 Subject: [PATCH 767/829] feat (Lattice/Covolume): Add results linking point counting in lattices and covolume (#12488) Co-authored-by: Xavier Roblot <46200072+xroblot@users.noreply.github.com> --- Mathlib/Algebra/Module/ZLattice/Basic.lean | 17 ++ Mathlib/Algebra/Module/ZLattice/Covolume.lean | 238 +++++++++++++++++- Mathlib/LinearAlgebra/Determinant.lean | 15 ++ 3 files changed, 267 insertions(+), 3 deletions(-) diff --git a/Mathlib/Algebra/Module/ZLattice/Basic.lean b/Mathlib/Algebra/Module/ZLattice/Basic.lean index b8b6b0b57e2f6..05499a60dcb19 100644 --- a/Mathlib/Algebra/Module/ZLattice/Basic.lean +++ b/Mathlib/Algebra/Module/ZLattice/Basic.lean @@ -708,4 +708,21 @@ theorem Basis.ofZLatticeComap_repr_apply (e : F ≃ₗ[K] E) {ι : Type*} (b : B end comap +section NormedLinearOrderedField_comap + +variable (K : Type*) [NormedLinearOrderedField K] [HasSolidNorm K] [FloorRing K] +variable {E : Type*} [NormedAddCommGroup E] [NormedSpace K E] [FiniteDimensional K E] + [ProperSpace E] +variable {F : Type*} [NormedAddCommGroup F] [NormedSpace K F] [FiniteDimensional K F] + [ProperSpace F] +variable (L : Submodule ℤ E) [DiscreteTopology L] [IsZLattice K L] + +theorem Basis.ofZLatticeBasis_comap (e : F ≃L[K] E) {ι : Type*} (b : Basis ι ℤ L) : + (b.ofZLatticeComap K L e.toLinearEquiv).ofZLatticeBasis K (ZLattice.comap K L e.toLinearMap) = + (b.ofZLatticeBasis K L).map e.symm.toLinearEquiv := by + ext + simp + +end NormedLinearOrderedField_comap + end ZLattice diff --git a/Mathlib/Algebra/Module/ZLattice/Covolume.lean b/Mathlib/Algebra/Module/ZLattice/Covolume.lean index 5f076d331f629..7886d9101873e 100644 --- a/Mathlib/Algebra/Module/ZLattice/Covolume.lean +++ b/Mathlib/Algebra/Module/ZLattice/Covolume.lean @@ -3,7 +3,8 @@ Copyright (c) 2024 Xavier Roblot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Xavier Roblot -/ -import Mathlib.Algebra.Module.ZLattice.Basic +import Mathlib.Analysis.BoxIntegral.UnitPartition +import Mathlib.MeasureTheory.Measure.Haar.InnerProductSpace /-! # Covolume of ℤ-lattices @@ -23,6 +24,26 @@ choice of the fundamental domain of `L`. * `ZLattice.covolume_eq_det`: if `L` is a lattice in `ℝ^n`, then its covolume is the absolute value of the determinant of any `ℤ`-basis of `L`. +* `ZLattice.covolume.tendsto_card_div_pow`: Let `s` be a bounded measurable set of `ι → ℝ`, then +the number of points in `s ∩ n⁻¹ • L` divided by `n ^ card ι` tends to `volume s / covolume L` +when `n : ℕ` tends to infinity. See also `ZLattice.covolume.tendsto_card_div_pow'` for a version +for `InnerProductSpace ℝ E` and `ZLattice.covolume.tendsto_card_div_pow''` for the general version. + +* `ZLattice.covolume.tendsto_card_le_div`: Let `X` be a cone in `ι → ℝ` and let `F : (ι → ℝ) → ℝ` +be a function such that `F (c • x) = c ^ card ι * F x`. Then the number of points `x ∈ X` such that +`F x ≤ c` divided by `c` tends to `volume {x ∈ X | F x ≤ 1} / covolume L` when `c : ℝ` tends to +infinity. See also `ZLattice.covolume.tendsto_card_le_div'` for a version for +`InnerProductSpace ℝ E` and `ZLattice.covolume.tendsto_card_le_div''` for the general version. + +## Naming convention + +Some results are true in the case where the ambient finite dimensional real vector space is the +pi-space `ι → ℝ` and in the case where it is an `InnerProductSpace`. We use the following +convention: the plain name is for the pi case, for eg. `volume_image_eq_volume_div_covolume`. For +the same result in the `InnerProductSpace` case, we add a `prime`, for eg. +`volume_image_eq_volume_div_covolume'`. When the same result exists in the +general case, we had two primes, eg. `covolume.tendsto_card_div_pow''`. + -/ noncomputable section @@ -42,7 +63,7 @@ def covolume (μ : Measure E := by volume_tac) : ℝ := (addCovolume L E μ).toR end General -section Real +section Basic variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [FiniteDimensional ℝ E] variable [MeasurableSpace E] [BorelSpace E] @@ -96,6 +117,217 @@ theorem covolume_eq_det {ι : Type*} [Fintype ι] [DecidableEq ι] (L : Submodul ext1 exact b.ofZLatticeBasis_apply ℝ L _ -end Real +theorem covolume_eq_det_inv {ι : Type*} [Fintype ι] [DecidableEq ι] (L : Submodule ℤ (ι → ℝ)) + [DiscreteTopology L] [IsZLattice ℝ L] (b : Basis ι ℤ L) : + covolume L = |(LinearEquiv.det (b.ofZLatticeBasis ℝ L).equivFun : ℝ)|⁻¹ := by + rw [covolume_eq_det L b, ← Pi.basisFun_det_apply, show (((↑) : L → _) ∘ ⇑b) = + (b.ofZLatticeBasis ℝ) by ext; simp, ← Basis.det_inv, ← abs_inv, Units.val_inv_eq_inv_val, + IsUnit.unit_spec, ← Basis.det_basis, LinearEquiv.coe_det] + rfl + +theorem volume_image_eq_volume_div_covolume {ι : Type*} [Fintype ι] [DecidableEq ι] + (L : Submodule ℤ (ι → ℝ)) [DiscreteTopology L] [IsZLattice ℝ L] (b : Basis ι ℤ L) + {s : Set (ι → ℝ)} : + volume ((b.ofZLatticeBasis ℝ L).equivFun '' s) = volume s / ENNReal.ofReal (covolume L) := by + rw [LinearEquiv.image_eq_preimage, Measure.addHaar_preimage_linearEquiv, LinearEquiv.symm_symm, + covolume_eq_det_inv L b, ENNReal.div_eq_inv_mul, ENNReal.ofReal_inv_of_pos + (abs_pos.mpr (LinearEquiv.det _).ne_zero), inv_inv, LinearEquiv.coe_det] + +/-- A more general version of `ZLattice.volume_image_eq_volume_div_covolume`; +see the `Naming conventions` section in the introduction. -/ +theorem volume_image_eq_volume_div_covolume' {E : Type*} [NormedAddCommGroup E] + [InnerProductSpace ℝ E] [FiniteDimensional ℝ E] [MeasurableSpace E] [BorelSpace E] + (L : Submodule ℤ E) [DiscreteTopology L] [IsZLattice ℝ L] {ι : Type*} [Fintype ι] + (b : Basis ι ℤ L) {s : Set E} (hs : NullMeasurableSet s) : + volume ((b.ofZLatticeBasis ℝ).equivFun '' s) = volume s / ENNReal.ofReal (covolume L) := by + classical + let e : Fin (finrank ℝ E) ≃ ι := + Fintype.equivOfCardEq (by rw [Fintype.card_fin, finrank_eq_card_basis (b.ofZLatticeBasis ℝ)]) + let f := (EuclideanSpace.equiv ι ℝ).symm.trans + ((stdOrthonormalBasis ℝ E).reindex e).repr.toContinuousLinearEquiv.symm + have hf : MeasurePreserving f := + ((stdOrthonormalBasis ℝ E).reindex e).measurePreserving_repr_symm.comp + (EuclideanSpace.volume_preserving_measurableEquiv ι).symm + rw [← hf.measure_preimage hs, ← (covolume_comap L volume volume hf), + ← volume_image_eq_volume_div_covolume (ZLattice.comap ℝ L f.toLinearMap) + (b.ofZLatticeComap ℝ L f.toLinearEquiv), Basis.ofZLatticeBasis_comap, + ← f.image_symm_eq_preimage, ← Set.image_comp] + simp only [Basis.equivFun_apply, ContinuousLinearEquiv.symm_toLinearEquiv, Basis.map_equivFun, + LinearEquiv.symm_symm, Function.comp_apply, LinearEquiv.trans_apply, + ContinuousLinearEquiv.coe_toLinearEquiv, ContinuousLinearEquiv.apply_symm_apply] + +end Basic + +namespace covolume + +section General + +open Filter Fintype Pointwise Topology BoxIntegral Bornology + +variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] +variable {L : Submodule ℤ E} [DiscreteTopology L] [IsZLattice ℝ L] +variable {ι : Type*} [Fintype ι] (b : Basis ι ℤ L) + +/-- A version of `ZLattice.covolume.tendsto_card_div_pow` for the general case; +see the `Naming convention` section in the introduction. -/ +theorem tendsto_card_div_pow'' [FiniteDimensional ℝ E] [MeasurableSpace E] [BorelSpace E] + {s : Set E} (hs₁ : IsBounded s) (hs₂ : MeasurableSet s) + (hs₃ : volume (frontier ((b.ofZLatticeBasis ℝ).equivFun '' s)) = 0): + Tendsto (fun n : ℕ ↦ (Nat.card (s ∩ (n : ℝ)⁻¹ • L : Set E) : ℝ) / n ^ card ι) + atTop (𝓝 (volume ((b.ofZLatticeBasis ℝ).equivFun '' s)).toReal) := by + refine Tendsto.congr' ?_ + (tendsto_card_div_pow_atTop_volume ((b.ofZLatticeBasis ℝ).equivFun '' s) ?_ ?_ hs₃) + · filter_upwards [eventually_gt_atTop 0] with n hn + congr + refine Nat.card_congr <| ((b.ofZLatticeBasis ℝ).equivFun.toEquiv.subtypeEquiv fun x ↦ ?_).symm + simp_rw [Set.mem_inter_iff, ← b.ofZLatticeBasis_span ℝ, LinearEquiv.coe_toEquiv, + Basis.equivFun_apply, Set.mem_image, DFunLike.coe_fn_eq, EmbeddingLike.apply_eq_iff_eq, + exists_eq_right, and_congr_right_iff, Set.mem_inv_smul_set_iff₀ + (mod_cast hn.ne' : (n : ℝ) ≠ 0), ← Finsupp.coe_smul, ← LinearEquiv.map_smul, SetLike.mem_coe, + Basis.mem_span_iff_repr_mem, Pi.basisFun_repr, implies_true] + · rw [← NormedSpace.isVonNBounded_iff ℝ] at hs₁ ⊢ + exact Bornology.IsVonNBounded.image hs₁ ((b.ofZLatticeBasis ℝ).equivFunL : E →L[ℝ] ι → ℝ) + · exact (b.ofZLatticeBasis ℝ).equivFunL.toHomeomorph.toMeasurableEquiv.measurableSet_image.mpr hs₂ + +private theorem tendsto_card_le_div''_aux {X : Set E} (hX : ∀ ⦃x⦄ ⦃r:ℝ⦄, x ∈ X → 0 < r → r • x ∈ X) + {F : E → ℝ} (hF₁ : ∀ x ⦃r : ℝ⦄, 0 ≤ r → F (r • x) = r ^ card ι * (F x)) {c : ℝ} (hc : 0 < c) : + c • {x ∈ X | F x ≤ 1} = {x ∈ X | F x ≤ c ^ card ι} := by + ext x + simp_rw [Set.mem_smul_set_iff_inv_smul_mem₀ hc.ne', Set.mem_setOf_eq, hF₁ _ + (inv_pos_of_pos hc).le, inv_pow, inv_mul_le_iff₀ (pow_pos hc _), mul_one, and_congr_left_iff] + exact fun _ ↦ ⟨fun h ↦ (smul_inv_smul₀ hc.ne' x) ▸ hX h hc, fun h ↦ hX h (inv_pos_of_pos hc)⟩ + +/-- A version of `ZLattice.covolume.tendsto_card_le_div` for the general case; +see the `Naming conventions` section in the introduction. -/ +theorem tendsto_card_le_div'' [FiniteDimensional ℝ E] [MeasurableSpace E] [BorelSpace E] + [Nonempty ι] {X : Set E} (hX : ∀ ⦃x⦄ ⦃r : ℝ⦄, x ∈ X → 0 < r → r • x ∈ X) + {F : E → ℝ} (h₁ : ∀ x ⦃r : ℝ⦄, 0 ≤ r → F (r • x) = r ^ card ι * (F x)) + (h₂ : IsBounded {x ∈ X | F x ≤ 1}) (h₃ : MeasurableSet {x ∈ X | F x ≤ 1}) + (h₄ : volume (frontier ((b.ofZLatticeBasis ℝ L).equivFun '' {x | x ∈ X ∧ F x ≤ 1})) = 0) : + Tendsto (fun c : ℝ ↦ + Nat.card ({x ∈ X | F x ≤ c} ∩ L : Set E) / (c : ℝ)) + atTop (𝓝 (volume ((b.ofZLatticeBasis ℝ).equivFun '' {x ∈ X | F x ≤ 1})).toReal) := by + + refine Tendsto.congr' ?_ <| (tendsto_card_div_pow_atTop_volume' + ((b.ofZLatticeBasis ℝ).equivFun '' {x ∈ X | F x ≤ 1}) ?_ ?_ h₄ fun x y hx hy ↦ ?_).comp + (tendsto_rpow_atTop <| inv_pos.mpr + (Nat.cast_pos.mpr card_pos) : Tendsto (fun x ↦ x ^ (card ι : ℝ)⁻¹) atTop atTop) + · filter_upwards [eventually_gt_atTop 0] with c hc + have aux₁ : (card ι : ℝ) ≠ 0 := Nat.cast_ne_zero.mpr card_ne_zero + have aux₂ : 0 < c ^ (card ι : ℝ)⁻¹ := Real.rpow_pos_of_pos hc _ + have aux₃ : (c ^ (card ι : ℝ)⁻¹)⁻¹ ≠ 0 := inv_ne_zero aux₂.ne' + have aux₄ : c ^ (-(card ι : ℝ)⁻¹) ≠ 0 := (Real.rpow_pos_of_pos hc _).ne' + obtain ⟨hc₁, hc₂⟩ := lt_iff_le_and_ne.mp hc + rw [Function.comp_apply, ← Real.rpow_natCast, Real.rpow_inv_rpow hc₁ aux₁, eq_comm] + congr + refine Nat.card_congr <| Equiv.subtypeEquiv ((b.ofZLatticeBasis ℝ).equivFun.toEquiv.trans + (Equiv.smulRight aux₄)) fun _ ↦ ?_ + rw [Set.mem_inter_iff, Set.mem_inter_iff, Equiv.trans_apply, LinearEquiv.coe_toEquiv, + Equiv.smulRight_apply, Real.rpow_neg hc₁, Set.smul_mem_smul_set_iff₀ aux₃, + ← Set.mem_smul_set_iff_inv_smul_mem₀ aux₂.ne', ← image_smul_set, + tendsto_card_le_div''_aux hX h₁ aux₂, ← Real.rpow_natCast, ← Real.rpow_mul hc₁, + inv_mul_cancel₀ aux₁, Real.rpow_one] + simp_rw [SetLike.mem_coe, Set.mem_image, EmbeddingLike.apply_eq_iff_eq, exists_eq_right, + and_congr_right_iff, ← b.ofZLatticeBasis_span ℝ, Basis.mem_span_iff_repr_mem, + Pi.basisFun_repr, Basis.equivFun_apply, implies_true] + · rw [← NormedSpace.isVonNBounded_iff ℝ] at h₂ ⊢ + exact Bornology.IsVonNBounded.image h₂ ((b.ofZLatticeBasis ℝ).equivFunL : E →L[ℝ] ι → ℝ) + · exact (b.ofZLatticeBasis ℝ).equivFunL.toHomeomorph.toMeasurableEquiv.measurableSet_image.mpr h₃ + · simp_rw [← image_smul_set] + apply Set.image_mono + rw [tendsto_card_le_div''_aux hX h₁ hx, + tendsto_card_le_div''_aux hX h₁ (lt_of_lt_of_le hx hy)] + exact fun a ⟨ha₁, ha₂⟩ ↦ ⟨ha₁, le_trans ha₂ <| pow_le_pow_left₀ (le_of_lt hx) hy _⟩ + +end General + +section Pi + +open Filter Fintype Pointwise Topology Bornology + +private theorem frontier_equivFun {E : Type*} [AddCommGroup E] [Module ℝ E] {ι : Type*} [Fintype ι] + [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousSMul ℝ E] [T2Space E] + (b : Basis ι ℝ E) (s : Set E) : + frontier (b.equivFun '' s) = b.equivFun '' (frontier s) := by + rw [LinearEquiv.image_eq_preimage, LinearEquiv.image_eq_preimage] + exact (Homeomorph.preimage_frontier b.equivFunL.toHomeomorph.symm s).symm + +variable {ι : Type*} [Fintype ι] +variable (L : Submodule ℤ (ι → ℝ)) [DiscreteTopology L] [IsZLattice ℝ L] + +theorem tendsto_card_div_pow (b : Basis ι ℤ L) {s : Set (ι → ℝ)} (hs₁ : IsBounded s) + (hs₂ : MeasurableSet s) (hs₃ : volume (frontier s) = 0) : + Tendsto (fun n : ℕ ↦ (Nat.card (s ∩ (n : ℝ)⁻¹ • L : Set (ι → ℝ)) : ℝ) / n ^ card ι) + atTop (𝓝 ((volume s).toReal / covolume L)) := by + classical + convert tendsto_card_div_pow'' b hs₁ hs₂ ?_ + · rw [volume_image_eq_volume_div_covolume L b, ENNReal.toReal_div, + ENNReal.toReal_ofReal (covolume_pos L volume).le] + · rw [frontier_equivFun, volume_image_eq_volume_div_covolume, hs₃, ENNReal.zero_div] + +theorem tendsto_card_le_div {X : Set (ι → ℝ)} (hX : ∀ ⦃x⦄ ⦃r : ℝ⦄, x ∈ X → 0 < r → r • x ∈ X) + {F : (ι → ℝ) → ℝ} (h₁ : ∀ x ⦃r : ℝ⦄, 0 ≤ r → F (r • x) = r ^ card ι * (F x)) + (h₂ : IsBounded {x ∈ X | F x ≤ 1}) (h₃ : MeasurableSet {x ∈ X | F x ≤ 1}) + (h₄ : volume (frontier {x | x ∈ X ∧ F x ≤ 1}) = 0) [Nonempty ι] : + Tendsto (fun c : ℝ ↦ + Nat.card ({x ∈ X | F x ≤ c} ∩ L : Set (ι → ℝ)) / (c : ℝ)) + atTop (𝓝 ((volume {x ∈ X | F x ≤ 1}).toReal / covolume L)) := by + classical + let e : Free.ChooseBasisIndex ℤ ↥L ≃ ι := by + refine Fintype.equivOfCardEq ?_ + rw [← finrank_eq_card_chooseBasisIndex, ZLattice.rank ℝ, finrank_fintype_fun_eq_card] + let b := (Module.Free.chooseBasis ℤ L).reindex e + convert tendsto_card_le_div'' b hX h₁ h₂ h₃ ?_ + · rw [volume_image_eq_volume_div_covolume L b, ENNReal.toReal_div, + ENNReal.toReal_ofReal (covolume_pos L volume).le] + · rw [frontier_equivFun, volume_image_eq_volume_div_covolume, h₄, ENNReal.zero_div] + +end Pi + +section InnerProductSpace + +open Filter Pointwise Topology Bornology + +variable {E : Type*} [NormedAddCommGroup E] [InnerProductSpace ℝ E] [FiniteDimensional ℝ E] + [MeasurableSpace E] [BorelSpace E] +variable (L : Submodule ℤ E) [DiscreteTopology L] [IsZLattice ℝ L] + +/-- A version of `ZLattice.covolume.tendsto_card_div_pow` for the `InnerProductSpace` case; +see the `Naming convention` section in the introduction. -/ +theorem tendsto_card_div_pow' {s : Set E} (hs₁ : IsBounded s) (hs₂ : MeasurableSet s) + (hs₃ : volume (frontier s) = 0) : + Tendsto (fun n : ℕ ↦ (Nat.card (s ∩ (n : ℝ)⁻¹ • L : Set E) : ℝ) / n ^ finrank ℝ E) + atTop (𝓝 ((volume s).toReal / covolume L)) := by + let b := Module.Free.chooseBasis ℤ L + convert tendsto_card_div_pow'' b hs₁ hs₂ ?_ + · rw [← finrank_eq_card_chooseBasisIndex, ZLattice.rank ℝ L] + · rw [volume_image_eq_volume_div_covolume' L b hs₂.nullMeasurableSet, ENNReal.toReal_div, + ENNReal.toReal_ofReal (covolume_pos L volume).le] + · rw [frontier_equivFun, volume_image_eq_volume_div_covolume', hs₃, ENNReal.zero_div] + exact NullMeasurableSet.of_null hs₃ + +/-- A version of `ZLattice.covolume.tendsto_card_le_div` for the `InnerProductSpace` case; +see the `Naming convention` section in the introduction. -/ +theorem tendsto_card_le_div' [Nontrivial E] {X : Set E} {F : E → ℝ} + (hX : ∀ ⦃x⦄ ⦃r : ℝ⦄, x ∈ X → 0 < r → r • x ∈ X) + (h₁ : ∀ x ⦃r : ℝ⦄, 0 ≤ r → F (r • x) = r ^ finrank ℝ E * (F x)) + (h₂ : IsBounded {x ∈ X | F x ≤ 1}) (h₃ : MeasurableSet {x ∈ X | F x ≤ 1}) + (h₄ : volume (frontier {x ∈ X | F x ≤ 1}) = 0) : + Tendsto (fun c : ℝ ↦ + Nat.card ({x ∈ X | F x ≤ c} ∩ L : Set E) / (c : ℝ)) + atTop (𝓝 ((volume {x ∈ X | F x ≤ 1}).toReal / covolume L)) := by + let b := Module.Free.chooseBasis ℤ L + convert tendsto_card_le_div'' b hX ?_ h₂ h₃ ?_ + · rw [volume_image_eq_volume_div_covolume' L b h₃.nullMeasurableSet, ENNReal.toReal_div, + ENNReal.toReal_ofReal (covolume_pos L volume).le] + · have : Nontrivial L := nontrivial_of_finrank_pos <| (ZLattice.rank ℝ L).symm ▸ finrank_pos + infer_instance + · rwa [← finrank_eq_card_chooseBasisIndex, ZLattice.rank ℝ L] + · rw [frontier_equivFun, volume_image_eq_volume_div_covolume', h₄, ENNReal.zero_div] + exact NullMeasurableSet.of_null h₄ + +end InnerProductSpace + +end covolume end ZLattice diff --git a/Mathlib/LinearAlgebra/Determinant.lean b/Mathlib/LinearAlgebra/Determinant.lean index 6a4b2827498d0..e707a63d57cbb 100644 --- a/Mathlib/LinearAlgebra/Determinant.lean +++ b/Mathlib/LinearAlgebra/Determinant.lean @@ -553,6 +553,16 @@ theorem Basis.det_comp_basis [Module A M'] (b : Basis ι A M) (b' : Basis ι A M congr 1; ext i j rw [Basis.toMatrix_apply, LinearMap.toMatrix_apply, Function.comp_apply] +@[simp] +theorem Basis.det_basis (b : Basis ι A M) (b' : Basis ι A M) : + LinearMap.det (b'.equiv b (Equiv.refl ι)).toLinearMap = b'.det b := + (b.det_comp_basis b' (LinearMap.id)).symm + +theorem Basis.det_inv (b : Basis ι A M) (b' : Basis ι A M) : + (b.isUnit_det b').unit⁻¹ = b'.det b := by + rw [← Units.mul_eq_one_iff_inv_eq, IsUnit.unit_spec, ← Basis.det_basis, ← Basis.det_basis] + exact LinearEquiv.det_mul_det_symm _ + theorem Basis.det_reindex {ι' : Type*} [Fintype ι'] [DecidableEq ι'] (b : Basis ι R M) (v : ι' → M) (e : ι ≃ ι') : (b.reindex e).det v = b.det (v ∘ e) := by rw [Basis.det_apply, Basis.toMatrix_reindex', det_reindexAlgEquiv, Basis.det_apply] @@ -579,6 +589,11 @@ theorem Pi.basisFun_det : (Pi.basisFun R ι).det = Matrix.detRowAlternating := b ext M rw [Basis.det_apply, Basis.coePiBasisFun.toMatrix_eq_transpose, det_transpose] +theorem Pi.basisFun_det_apply (v : ι → ι → R) : + (Pi.basisFun R ι).det v = (Matrix.of v).det := by + rw [Pi.basisFun_det] + rfl + /-- If we fix a background basis `e`, then for any other basis `v`, we can characterise the coordinates provided by `v` in terms of determinants relative to `e`. -/ theorem Basis.det_smul_mk_coord_eq_det_update {v : ι → M} (hli : LinearIndependent R v) From e282c781c5137c4a6cb1dd02ac0a2c7359c1ea04 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 18 Dec 2024 16:02:26 +0000 Subject: [PATCH 768/829] =?UTF-8?q?feat(ENNReal):=20`a=20*=20b=20*=20b?= =?UTF-8?q?=E2=81=BB=C2=B9=20=3D=20a`=20(#19870)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From FLT --- Archive/Wiedijk100Theorems/BallotProblem.lean | 2 +- Mathlib/Analysis/Oscillation.lean | 2 +- Mathlib/Data/ENNReal/Inv.lean | 84 +++++++++++++++++-- .../MeasureTheory/Covering/Besicovitch.lean | 2 +- .../Function/UniformIntegrable.lean | 2 +- Mathlib/MeasureTheory/Group/Prod.lean | 4 +- Mathlib/MeasureTheory/Integral/Average.lean | 2 +- .../Integral/VitaliCaratheodory.lean | 4 +- scripts/nolints_prime_decls.txt | 2 +- 9 files changed, 87 insertions(+), 17 deletions(-) diff --git a/Archive/Wiedijk100Theorems/BallotProblem.lean b/Archive/Wiedijk100Theorems/BallotProblem.lean index d11275c43b41a..1476d785b39a2 100644 --- a/Archive/Wiedijk100Theorems/BallotProblem.lean +++ b/Archive/Wiedijk100Theorems/BallotProblem.lean @@ -236,7 +236,7 @@ theorem first_vote_neg (p q : ℕ) (h : 0 < p + q) : {l : List ℤ | l.headI = 1}ᶜ (countedSequence_finite p q) (countedSequence_nonempty p q) rw [compl_compl, first_vote_pos _ _ h] at this 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] + ENNReal.mul_div_cancel, ENNReal.add_sub_cancel_left] all_goals simp_all [ENNReal.div_eq_top] theorem ballot_same (p : ℕ) : uniformOn (countedSequence (p + 1) (p + 1)) staysPositive = 0 := by diff --git a/Mathlib/Analysis/Oscillation.lean b/Mathlib/Analysis/Oscillation.lean index a9e4ec628acf0..86dd22a83081a 100644 --- a/Mathlib/Analysis/Oscillation.lean +++ b/Mathlib/Analysis/Oscillation.lean @@ -56,7 +56,7 @@ theorem oscillationWithin_eq_zero [TopologicalSpace E] {f : E → F} {D : Set E} rw [zero_add] have : ball (f x) (ε / 2) ∈ (𝓝[D] x).map f := hf <| ball_mem_nhds _ (by simp [ne_of_gt hε]) refine (biInf_le diam this).trans (le_of_le_of_eq diam_ball ?_) - exact (ENNReal.mul_div_cancel' (by norm_num) (by norm_num)) + exact (ENNReal.mul_div_cancel (by norm_num) (by norm_num)) end ContinuousWithinAt diff --git a/Mathlib/Data/ENNReal/Inv.lean b/Mathlib/Data/ENNReal/Inv.lean index 0a17941605844..066858d61ecb6 100644 --- a/Mathlib/Data/ENNReal/Inv.lean +++ b/Mathlib/Data/ENNReal/Inv.lean @@ -91,11 +91,81 @@ protected theorem mul_inv_cancel (h0 : a ≠ 0) (ht : a ≠ ∞) : a * a⁻¹ = protected theorem inv_mul_cancel (h0 : a ≠ 0) (ht : a ≠ ∞) : a⁻¹ * a = 1 := mul_comm a a⁻¹ ▸ ENNReal.mul_inv_cancel h0 ht -protected theorem div_mul_cancel (h0 : a ≠ 0) (hI : a ≠ ∞) : b / a * a = b := by - rw [div_eq_mul_inv, mul_assoc, ENNReal.inv_mul_cancel h0 hI, mul_one] - -protected theorem mul_div_cancel' (h0 : a ≠ 0) (hI : a ≠ ∞) : a * (b / a) = b := by - rw [mul_comm, ENNReal.div_mul_cancel h0 hI] +/-- See `ENNReal.inv_mul_cancel_left` for a simpler version assuming `a ≠ 0`, `a ≠ ∞`. -/ +protected lemma inv_mul_cancel_left' (ha₀ : a = 0 → b = 0) (ha : a = ∞ → b = 0) : + a⁻¹ * (a * b) = b := by + obtain rfl | ha₀ := eq_or_ne a 0 + · simp_all + obtain rfl | ha := eq_or_ne a ⊤ + · simp_all + · simp [← mul_assoc, ENNReal.inv_mul_cancel, *] + +/-- See `ENNReal.inv_mul_cancel_left'` for a stronger version. -/ +protected lemma inv_mul_cancel_left (ha₀ : a ≠ 0) (ha : a ≠ ∞) : a⁻¹ * (a * b) = b := + ENNReal.inv_mul_cancel_left' (by simp [ha₀]) (by simp [ha]) + +/-- See `ENNReal.mul_inv_cancel_left` for a simpler version assuming `a ≠ 0`, `a ≠ ∞`. -/ +protected lemma mul_inv_cancel_left' (ha₀ : a = 0 → b = 0) (ha : a = ∞ → b = 0) : + a * (a⁻¹ * b) = b := by + obtain rfl | ha₀ := eq_or_ne a 0 + · simp_all + obtain rfl | ha := eq_or_ne a ⊤ + · simp_all + · simp [← mul_assoc, ENNReal.mul_inv_cancel, *] + +/-- See `ENNReal.mul_inv_cancel_left'` for a stronger version. -/ +protected lemma mul_inv_cancel_left (ha₀ : a ≠ 0) (ha : a ≠ ∞) : a * (a⁻¹ * b) = b := + ENNReal.mul_inv_cancel_left' (by simp [ha₀]) (by simp [ha]) + +/-- See `ENNReal.mul_inv_cancel_right` for a simpler version assuming `b ≠ 0`, `b ≠ ∞`. -/ +protected lemma mul_inv_cancel_right' (hb₀ : b = 0 → a = 0) (hb : b = ∞ → a = 0) : + a * b * b⁻¹ = a := by + obtain rfl | hb₀ := eq_or_ne b 0 + · simp_all + obtain rfl | hb := eq_or_ne b ⊤ + · simp_all + · simp [mul_assoc, ENNReal.mul_inv_cancel, *] + +/-- See `ENNReal.mul_inv_cancel_right'` for a stronger version. -/ +protected lemma mul_inv_cancel_right (hb₀ : b ≠ 0) (hb : b ≠ ∞) : a * b * b⁻¹ = a := + ENNReal.mul_inv_cancel_right' (by simp [hb₀]) (by simp [hb]) + +/-- See `ENNReal.inv_mul_cancel_right` for a simpler version assuming `b ≠ 0`, `b ≠ ∞`. -/ +protected lemma inv_mul_cancel_right' (hb₀ : b = 0 → a = 0) (hb : b = ∞ → a = 0) : + a * b⁻¹ * b = a := by + obtain rfl | hb₀ := eq_or_ne b 0 + · simp_all + obtain rfl | hb := eq_or_ne b ⊤ + · simp_all + · simp [mul_assoc, ENNReal.inv_mul_cancel, *] + +/-- See `ENNReal.inv_mul_cancel_right'` for a stronger version. -/ +protected lemma inv_mul_cancel_right (hb₀ : b ≠ 0) (hb : b ≠ ∞) : a * b⁻¹ * b = a := + ENNReal.inv_mul_cancel_right' (by simp [hb₀]) (by simp [hb]) + +/-- See `ENNReal.mul_div_cancel_right` for a simpler version assuming `b ≠ 0`, `b ≠ ∞`. -/ +protected lemma mul_div_cancel_right' (hb₀ : b = 0 → a = 0) (hb : b = ∞ → a = 0) : + a * b / b = a := ENNReal.mul_inv_cancel_right' hb₀ hb + +/-- See `ENNReal.mul_div_cancel_right'` for a stronger version. -/ +protected lemma mul_div_cancel_right (hb₀ : b ≠ 0) (hb : b ≠ ∞) : a * b / b = a := + ENNReal.mul_div_cancel_right' (by simp [hb₀]) (by simp [hb]) + +/-- See `ENNReal.div_mul_cancel` for a simpler version assuming `a ≠ 0`, `a ≠ ∞`. -/ +protected lemma div_mul_cancel' (ha₀ : a = 0 → b = 0) (ha : a = ∞ → b = 0) : b / a * a = b := + ENNReal.inv_mul_cancel_right' ha₀ ha + +/-- See `ENNReal.div_mul_cancel'` for a stronger version. -/ +protected lemma div_mul_cancel (ha₀ : a ≠ 0) (ha : a ≠ ∞) : b / a * a = b := + ENNReal.div_mul_cancel' (by simp [ha₀]) (by simp [ha]) + +/-- See `ENNReal.mul_div_cancel` for a simpler version assuming `a ≠ 0`, `a ≠ ∞`. -/ +protected lemma mul_div_cancel' (ha₀ : a = 0 → b = 0) (ha : a = ∞ → b = 0) : a * (b / a) = b := by + rw [mul_comm, ENNReal.div_mul_cancel' ha₀ ha] + +/-- See `ENNReal.mul_div_cancel'` for a stronger version. -/ +protected lemma mul_div_cancel (ha₀ : a ≠ 0) (ha : a ≠ ∞) : a * (b / a) = b := + ENNReal.mul_div_cancel' (by simp [ha₀]) (by simp [ha]) -- Porting note: `simp only [div_eq_mul_inv, mul_comm, mul_assoc]` doesn't work in the following two protected theorem mul_comm_div : a / b * c = a * (c / b) := by @@ -398,8 +468,8 @@ theorem mul_div_le : a * (b / a) ≤ b := mul_le_of_le_div' le_rfl theorem eq_div_iff (ha : a ≠ 0) (ha' : a ≠ ∞) : b = c / a ↔ a * b = c := - ⟨fun h => by rw [h, ENNReal.mul_div_cancel' ha ha'], fun h => by - rw [← h, mul_div_assoc, ENNReal.mul_div_cancel' ha ha']⟩ + ⟨fun h => by rw [h, ENNReal.mul_div_cancel ha ha'], fun h => by + rw [← h, mul_div_assoc, ENNReal.mul_div_cancel ha ha']⟩ protected theorem div_eq_div_iff (ha : a ≠ 0) (ha' : a ≠ ∞) (hb : b ≠ 0) (hb' : b ≠ ∞) : c / b = d / a ↔ a * c = b * d := by diff --git a/Mathlib/MeasureTheory/Covering/Besicovitch.lean b/Mathlib/MeasureTheory/Covering/Besicovitch.lean index 8d6744400fbbc..f38954c20932f 100644 --- a/Mathlib/MeasureTheory/Covering/Besicovitch.lean +++ b/Mathlib/MeasureTheory/Covering/Besicovitch.lean @@ -575,7 +575,7 @@ theorem exist_finset_disjoint_balls_large_measure (μ : Measure α) [IsFiniteMea calc ∑ _i : Fin N, μ s / N = μ s := by simp only [Finset.card_fin, Finset.sum_const, nsmul_eq_mul] - rw [ENNReal.mul_div_cancel'] + rw [ENNReal.mul_div_cancel] · simp only [Npos, Ne, Nat.cast_eq_zero, not_false_iff] · exact ENNReal.natCast_ne_top _ _ ≤ ∑ i, μ (s ∩ v i) := by diff --git a/Mathlib/MeasureTheory/Function/UniformIntegrable.lean b/Mathlib/MeasureTheory/Function/UniformIntegrable.lean index c7c8a68fdf2d0..172756a4d9dbd 100644 --- a/Mathlib/MeasureTheory/Function/UniformIntegrable.lean +++ b/Mathlib/MeasureTheory/Function/UniformIntegrable.lean @@ -402,7 +402,7 @@ theorem Memℒp.eLpNorm_indicator_le_of_meas (hp_one : 1 ≤ p) (hp_top : p ≠ obtain ⟨δ, hδpos, hδ⟩ := hf.eLpNorm_indicator_le' hp_one hp_top hmeas (half_pos hε) refine ⟨δ, hδpos, fun s hs hμs => le_trans (hδ s hs hμs) ?_⟩ rw [ENNReal.ofReal_div_of_pos zero_lt_two, (by norm_num : ENNReal.ofReal 2 = 2), - ENNReal.mul_div_cancel'] <;> + ENNReal.mul_div_cancel] <;> norm_num @[deprecated (since := "2024-07-27")] diff --git a/Mathlib/MeasureTheory/Group/Prod.lean b/Mathlib/MeasureTheory/Group/Prod.lean index db452a55297de..cf0e2cb34ba04 100644 --- a/Mathlib/MeasureTheory/Group/Prod.lean +++ b/Mathlib/MeasureTheory/Group/Prod.lean @@ -305,7 +305,7 @@ theorem measure_lintegral_div_measure (sm : MeasurableSet s) (h2s : ν' s ≠ 0) simp_rw [measure_mul_lintegral_eq μ' ν' sm g hg, g, inv_inv] refine lintegral_congr_ae ?_ refine (ae_measure_preimage_mul_right_lt_top_of_ne_zero μ' ν' h2s h3s).mono fun x hx => ?_ - simp_rw [ENNReal.mul_div_cancel' (measure_mul_right_ne_zero ν' h2s _) hx.ne] + simp_rw [ENNReal.mul_div_cancel (measure_mul_right_ne_zero ν' h2s _) hx.ne] @[to_additive] theorem measure_mul_measure_eq (s t : Set G) (h2s : ν' s ≠ 0) (h3s : ν' s ≠ ∞) : @@ -330,7 +330,7 @@ theorem measure_eq_div_smul (h2s : ν' s ≠ 0) (h3s : ν' s ≠ ∞) : μ' = (μ' s / ν' s) • ν' := by ext1 t - rw [smul_apply, smul_eq_mul, mul_comm, ← mul_div_assoc, mul_comm, - measure_mul_measure_eq μ' ν' s t h2s h3s, mul_div_assoc, ENNReal.mul_div_cancel' h2s h3s] + measure_mul_measure_eq μ' ν' s t h2s h3s, mul_div_assoc, ENNReal.mul_div_cancel h2s h3s] end SigmaFinite diff --git a/Mathlib/MeasureTheory/Integral/Average.lean b/Mathlib/MeasureTheory/Integral/Average.lean index 7f6337976b25d..a315a1cda4c01 100644 --- a/Mathlib/MeasureTheory/Integral/Average.lean +++ b/Mathlib/MeasureTheory/Integral/Average.lean @@ -117,7 +117,7 @@ theorem measure_mul_laverage [IsFiniteMeasure μ] (f : α → ℝ≥0∞) : μ univ * ⨍⁻ x, f x ∂μ = ∫⁻ x, f x ∂μ := by rcases eq_or_ne μ 0 with hμ | hμ · rw [hμ, lintegral_zero_measure, laverage_zero_measure, mul_zero] - · rw [laverage_eq, ENNReal.mul_div_cancel' (measure_univ_ne_zero.2 hμ) (measure_ne_top _ _)] + · rw [laverage_eq, ENNReal.mul_div_cancel (measure_univ_ne_zero.2 hμ) (measure_ne_top _ _)] theorem setLaverage_eq (f : α → ℝ≥0∞) (s : Set α) : ⨍⁻ x in s, f x ∂μ = (∫⁻ x in s, f x ∂μ) / μ s := by rw [laverage_eq, restrict_apply_univ] diff --git a/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean b/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean index 4ebbecf29aa5d..b19b0ab5ee009 100644 --- a/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean +++ b/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean @@ -135,7 +135,7 @@ theorem SimpleFunc.exists_le_lowerSemicontinuous_lintegral_ge (f : α →ₛ ℝ (c : ℝ≥0∞) * μ u ≤ c * (μ s + ε / c) := mul_le_mul_left' μu.le _ _ = c * μ s + ε := by simp_rw [mul_add] - rw [ENNReal.mul_div_cancel' _ ENNReal.coe_ne_top] + rw [ENNReal.mul_div_cancel _ ENNReal.coe_ne_top] simpa using hc · rcases h₁ (ENNReal.half_pos ε0).ne' with ⟨g₁, f₁_le_g₁, g₁cont, g₁int⟩ @@ -351,7 +351,7 @@ theorem SimpleFunc.exists_upperSemicontinuous_le_lintegral_le (f : α →ₛ ℝ (c : ℝ≥0∞) * μ s ≤ c * (μ F + ε / c) := mul_le_mul_left' μF.le _ _ = c * μ F + ε := by simp_rw [mul_add] - rw [ENNReal.mul_div_cancel' _ ENNReal.coe_ne_top] + rw [ENNReal.mul_div_cancel _ ENNReal.coe_ne_top] simpa using hc · have A : ((∫⁻ x : α, f₁ x ∂μ) + ∫⁻ x : α, f₂ x ∂μ) ≠ ⊤ := by rwa [← lintegral_add_left f₁.measurable.coe_nnreal_ennreal] diff --git a/scripts/nolints_prime_decls.txt b/scripts/nolints_prime_decls.txt index 6924760aaad11..3304424a06b19 100644 --- a/scripts/nolints_prime_decls.txt +++ b/scripts/nolints_prime_decls.txt @@ -1147,7 +1147,7 @@ ENNReal.iInf_mul_right' ENNReal.inv_le_inv' ENNReal.inv_lt_inv' ENNReal.log_pos_real' -ENNReal.mul_div_cancel' +ENNReal.mul_div_cancel ENNReal.mul_le_of_le_div' ENNReal.mul_lt_mul_left' ENNReal.mul_lt_mul_right' From 78c7971c435461692c193bdc87b63ab700754e6f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 18 Dec 2024 16:02:29 +0000 Subject: [PATCH 769/829] feat(Data/Set/Lattice): `image2` lemmas (#20020) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Also add spaces after `⋃₀` for uniformity with `⋂₀`. From GrowthInGroups (LeanCamCombi) --- Mathlib/Data/Set/Lattice.lean | 108 ++++++++++++++++++++++------------ 1 file changed, 69 insertions(+), 39 deletions(-) diff --git a/Mathlib/Data/Set/Lattice.lean b/Mathlib/Data/Set/Lattice.lean index b77cd5d86cdec..19c7c55298ff4 100644 --- a/Mathlib/Data/Set/Lattice.lean +++ b/Mathlib/Data/Set/Lattice.lean @@ -56,7 +56,7 @@ open Function Set universe u -variable {α β γ : Type*} {ι ι' ι₂ : Sort*} {κ κ₁ κ₂ : ι → Sort*} {κ' : ι' → Sort*} +variable {α β γ δ : Type*} {ι ι' ι₂ : Sort*} {κ κ₁ κ₂ : ι → Sort*} {κ' : ι' → Sort*} namespace Set @@ -841,28 +841,28 @@ theorem iInter₂_union (s : ∀ i, κ i → Set α) (t : Set α) : (⋂ (i) (j), s i j) ∪ t = ⋂ (i) (j), s i j ∪ t := by simp_rw [iInter_union] theorem mem_sUnion_of_mem {x : α} {t : Set α} {S : Set (Set α)} (hx : x ∈ t) (ht : t ∈ S) : - x ∈ ⋃₀S := + x ∈ ⋃₀ S := ⟨t, ht, hx⟩ -- is this theorem really necessary? -theorem not_mem_of_not_mem_sUnion {x : α} {t : Set α} {S : Set (Set α)} (hx : x ∉ ⋃₀S) +theorem not_mem_of_not_mem_sUnion {x : α} {t : Set α} {S : Set (Set α)} (hx : x ∉ ⋃₀ S) (ht : t ∈ S) : x ∉ t := fun h => hx ⟨t, ht, h⟩ theorem sInter_subset_of_mem {S : Set (Set α)} {t : Set α} (tS : t ∈ S) : ⋂₀ S ⊆ t := sInf_le tS -theorem subset_sUnion_of_mem {S : Set (Set α)} {t : Set α} (tS : t ∈ S) : t ⊆ ⋃₀S := +theorem subset_sUnion_of_mem {S : Set (Set α)} {t : Set α} (tS : t ∈ S) : t ⊆ ⋃₀ S := le_sSup tS theorem subset_sUnion_of_subset {s : Set α} (t : Set (Set α)) (u : Set α) (h₁ : s ⊆ u) - (h₂ : u ∈ t) : s ⊆ ⋃₀t := + (h₂ : u ∈ t) : s ⊆ ⋃₀ t := Subset.trans h₁ (subset_sUnion_of_mem h₂) -theorem sUnion_subset {S : Set (Set α)} {t : Set α} (h : ∀ t' ∈ S, t' ⊆ t) : ⋃₀S ⊆ t := +theorem sUnion_subset {S : Set (Set α)} {t : Set α} (h : ∀ t' ∈ S, t' ⊆ t) : ⋃₀ S ⊆ t := sSup_le h @[simp] -theorem sUnion_subset_iff {s : Set (Set α)} {t : Set α} : ⋃₀s ⊆ t ↔ ∀ t' ∈ s, t' ⊆ t := +theorem sUnion_subset_iff {s : Set (Set α)} {t : Set α} : ⋃₀ s ⊆ t ↔ ∀ t' ∈ s, t' ⊆ t := sSup_le_iff /-- `sUnion` is monotone under taking a subset of each set. -/ @@ -884,7 +884,7 @@ theorem subset_sInter_iff {S : Set (Set α)} {t : Set α} : t ⊆ ⋂₀ S ↔ le_sInf_iff @[gcongr] -theorem sUnion_subset_sUnion {S T : Set (Set α)} (h : S ⊆ T) : ⋃₀S ⊆ ⋃₀T := +theorem sUnion_subset_sUnion {S T : Set (Set α)} (h : S ⊆ T) : ⋃₀ S ⊆ ⋃₀ T := sUnion_subset fun _ hs => subset_sUnion_of_mem (h hs) @[gcongr] @@ -892,7 +892,7 @@ theorem sInter_subset_sInter {S T : Set (Set α)} (h : S ⊆ T) : ⋂₀ T ⊆ subset_sInter fun _ hs => sInter_subset_of_mem (h hs) @[simp] -theorem sUnion_empty : ⋃₀∅ = (∅ : Set α) := +theorem sUnion_empty : ⋃₀ ∅ = (∅ : Set α) := sSup_empty @[simp] @@ -900,7 +900,7 @@ theorem sInter_empty : ⋂₀ ∅ = (univ : Set α) := sInf_empty @[simp] -theorem sUnion_singleton (s : Set α) : ⋃₀{s} = s := +theorem sUnion_singleton (s : Set α) : ⋃₀ {s} = s := sSup_singleton @[simp] @@ -908,7 +908,7 @@ theorem sInter_singleton (s : Set α) : ⋂₀ {s} = s := sInf_singleton @[simp] -theorem sUnion_eq_empty {S : Set (Set α)} : ⋃₀S = ∅ ↔ ∀ s ∈ S, s = ∅ := +theorem sUnion_eq_empty {S : Set (Set α)} : ⋃₀ S = ∅ ↔ ∀ s ∈ S, s = ∅ := sSup_eq_bot @[simp] @@ -939,24 +939,24 @@ theorem sUnion_mem_empty_univ {S : Set (Set α)} (h : S ⊆ {∅, univ}) : exact univ_subset_iff.1 <| subset_sUnion_of_mem hs @[simp] -theorem nonempty_sUnion {S : Set (Set α)} : (⋃₀S).Nonempty ↔ ∃ s ∈ S, Set.Nonempty s := by +theorem nonempty_sUnion {S : Set (Set α)} : (⋃₀ S).Nonempty ↔ ∃ s ∈ S, Set.Nonempty s := by simp [nonempty_iff_ne_empty] -theorem Nonempty.of_sUnion {s : Set (Set α)} (h : (⋃₀s).Nonempty) : s.Nonempty := +theorem Nonempty.of_sUnion {s : Set (Set α)} (h : (⋃₀ s).Nonempty) : s.Nonempty := let ⟨s, hs, _⟩ := nonempty_sUnion.1 h ⟨s, hs⟩ -theorem Nonempty.of_sUnion_eq_univ [Nonempty α] {s : Set (Set α)} (h : ⋃₀s = univ) : s.Nonempty := +theorem Nonempty.of_sUnion_eq_univ [Nonempty α] {s : Set (Set α)} (h : ⋃₀ s = univ) : s.Nonempty := Nonempty.of_sUnion <| h.symm ▸ univ_nonempty -theorem sUnion_union (S T : Set (Set α)) : ⋃₀(S ∪ T) = ⋃₀S ∪ ⋃₀T := +theorem sUnion_union (S T : Set (Set α)) : ⋃₀ (S ∪ T) = ⋃₀ S ∪ ⋃₀ T := sSup_union theorem sInter_union (S T : Set (Set α)) : ⋂₀ (S ∪ T) = ⋂₀ S ∩ ⋂₀ T := sInf_union @[simp] -theorem sUnion_insert (s : Set α) (T : Set (Set α)) : ⋃₀insert s T = s ∪ ⋃₀T := +theorem sUnion_insert (s : Set α) (T : Set (Set α)) : ⋃₀ insert s T = s ∪ ⋃₀ T := sSup_insert @[simp] @@ -964,29 +964,37 @@ theorem sInter_insert (s : Set α) (T : Set (Set α)) : ⋂₀ insert s T = s sInf_insert @[simp] -theorem sUnion_diff_singleton_empty (s : Set (Set α)) : ⋃₀(s \ {∅}) = ⋃₀s := +theorem sUnion_diff_singleton_empty (s : Set (Set α)) : ⋃₀ (s \ {∅}) = ⋃₀ s := sSup_diff_singleton_bot s @[simp] theorem sInter_diff_singleton_univ (s : Set (Set α)) : ⋂₀ (s \ {univ}) = ⋂₀ s := sInf_diff_singleton_top s -theorem sUnion_pair (s t : Set α) : ⋃₀{s, t} = s ∪ t := +theorem sUnion_pair (s t : Set α) : ⋃₀ {s, t} = s ∪ t := sSup_pair theorem sInter_pair (s t : Set α) : ⋂₀ {s, t} = s ∩ t := sInf_pair @[simp] -theorem sUnion_image (f : α → Set β) (s : Set α) : ⋃₀(f '' s) = ⋃ x ∈ s, f x := +theorem sUnion_image (f : α → Set β) (s : Set α) : ⋃₀ (f '' s) = ⋃ a ∈ s, f a := sSup_image @[simp] -theorem sInter_image (f : α → Set β) (s : Set α) : ⋂₀ (f '' s) = ⋂ x ∈ s, f x := +theorem sInter_image (f : α → Set β) (s : Set α) : ⋂₀ (f '' s) = ⋂ a ∈ s, f a := sInf_image @[simp] -theorem sUnion_range (f : ι → Set β) : ⋃₀range f = ⋃ x, f x := +lemma sUnion_image2 (f : α → β → Set γ) (s : Set α) (t : Set β) : + ⋃₀ (image2 f s t) = ⋃ (a ∈ s) (b ∈ t), f a b := sSup_image2 + +@[simp] +lemma sInter_image2 (f : α → β → Set γ) (s : Set α) (t : Set β) : + ⋂₀ (image2 f s t) = ⋂ (a ∈ s) (b ∈ t), f a b := sInf_image2 + +@[simp] +theorem sUnion_range (f : ι → Set β) : ⋃₀ range f = ⋃ x, f x := rfl @[simp] @@ -1001,7 +1009,7 @@ theorem iUnion₂_eq_univ_iff {s : ∀ i, κ i → Set α} : ⋃ (i) (j), s i j = univ ↔ ∀ a, ∃ i j, a ∈ s i j := by simp only [iUnion_eq_univ_iff, mem_iUnion] -theorem sUnion_eq_univ_iff {c : Set (Set α)} : ⋃₀c = univ ↔ ∀ a, ∃ b ∈ c, a ∈ b := by +theorem sUnion_eq_univ_iff {c : Set (Set α)} : ⋃₀ c = univ ↔ ∀ a, ∃ b ∈ c, a ∈ b := by simp only [eq_univ_iff_forall, mem_sUnion] -- classical @@ -1035,23 +1043,23 @@ theorem nonempty_sInter {c : Set (Set α)} : (⋂₀ c).Nonempty ↔ ∃ a, ∀ simp [nonempty_iff_ne_empty, sInter_eq_empty_iff] -- classical -theorem compl_sUnion (S : Set (Set α)) : (⋃₀S)ᶜ = ⋂₀ (compl '' S) := +theorem compl_sUnion (S : Set (Set α)) : (⋃₀ S)ᶜ = ⋂₀ (compl '' S) := ext fun x => by simp -- classical -theorem sUnion_eq_compl_sInter_compl (S : Set (Set α)) : ⋃₀S = (⋂₀ (compl '' S))ᶜ := by - rw [← compl_compl (⋃₀S), compl_sUnion] +theorem sUnion_eq_compl_sInter_compl (S : Set (Set α)) : ⋃₀ S = (⋂₀ (compl '' S))ᶜ := by + rw [← compl_compl (⋃₀ S), compl_sUnion] -- classical -theorem compl_sInter (S : Set (Set α)) : (⋂₀ S)ᶜ = ⋃₀(compl '' S) := by +theorem compl_sInter (S : Set (Set α)) : (⋂₀ S)ᶜ = ⋃₀ (compl '' S) := by rw [sUnion_eq_compl_sInter_compl, compl_compl_image] -- classical -theorem sInter_eq_compl_sUnion_compl (S : Set (Set α)) : ⋂₀ S = (⋃₀(compl '' S))ᶜ := by +theorem sInter_eq_compl_sUnion_compl (S : Set (Set α)) : ⋂₀ S = (⋃₀ (compl '' S))ᶜ := by rw [← compl_compl (⋂₀ S), compl_sInter] theorem inter_empty_of_inter_sUnion_empty {s t : Set α} {S : Set (Set α)} (hs : t ∈ S) - (h : s ∩ ⋃₀S = ∅) : s ∩ t = ∅ := + (h : s ∩ ⋃₀ S = ∅) : s ∩ t = ∅ := eq_empty_of_subset_empty <| by rw [← h]; exact inter_subset_inter_right _ (subset_sUnion_of_mem hs) @@ -1096,13 +1104,13 @@ theorem iUnion_of_singleton (α : Type*) : (⋃ x, {x} : Set α) = univ := by si theorem iUnion_of_singleton_coe (s : Set α) : ⋃ i : s, ({(i : α)} : Set α) = s := by simp -theorem sUnion_eq_biUnion {s : Set (Set α)} : ⋃₀s = ⋃ (i : Set α) (_ : i ∈ s), i := by +theorem sUnion_eq_biUnion {s : Set (Set α)} : ⋃₀ s = ⋃ (i : Set α) (_ : i ∈ s), i := by rw [← sUnion_image, image_id'] theorem sInter_eq_biInter {s : Set (Set α)} : ⋂₀ s = ⋂ (i : Set α) (_ : i ∈ s), i := by rw [← sInter_image, image_id'] -theorem sUnion_eq_iUnion {s : Set (Set α)} : ⋃₀s = ⋃ i : s, i := by +theorem sUnion_eq_iUnion {s : Set (Set α)} : ⋃₀ s = ⋃ i : s, i := by simp only [← sUnion_range, Subtype.range_coe] theorem sInter_eq_iInter {s : Set (Set α)} : ⋂₀ s = ⋂ i : s, i := by @@ -1127,7 +1135,7 @@ theorem sInter_union_sInter {S T : Set (Set α)} : sInf_sup_sInf theorem sUnion_inter_sUnion {s t : Set (Set α)} : - ⋃₀s ∩ ⋃₀t = ⋃ p ∈ s ×ˢ t, (p : Set α × Set α).1 ∩ p.2 := + ⋃₀ s ∩ ⋃₀ t = ⋃ p ∈ s ×ˢ t, (p : Set α × Set α).1 ∩ p.2 := sSup_inf_sSup theorem biUnion_iUnion (s : ι → Set α) (t : α → Set β) : @@ -1136,14 +1144,14 @@ theorem biUnion_iUnion (s : ι → Set α) (t : α → Set β) : theorem biInter_iUnion (s : ι → Set α) (t : α → Set β) : ⋂ x ∈ ⋃ i, s i, t x = ⋂ (i) (x ∈ s i), t x := by simp [@iInter_comm _ ι] -theorem sUnion_iUnion (s : ι → Set (Set α)) : ⋃₀⋃ i, s i = ⋃ i, ⋃₀s i := by +theorem sUnion_iUnion (s : ι → Set (Set α)) : ⋃₀ ⋃ i, s i = ⋃ i, ⋃₀ s i := by simp only [sUnion_eq_biUnion, biUnion_iUnion] theorem sInter_iUnion (s : ι → Set (Set α)) : ⋂₀ ⋃ i, s i = ⋂ i, ⋂₀ s i := by simp only [sInter_eq_biInter, biInter_iUnion] theorem iUnion_range_eq_sUnion {α β : Type*} (C : Set (Set α)) {f : ∀ s : C, β → (s : Type _)} - (hf : ∀ s : C, Surjective (f s)) : ⋃ y : β, range (fun s : C => (f s y).val) = ⋃₀C := by + (hf : ∀ s : C, Surjective (f s)) : ⋃ y : β, range (fun s : C => (f s y).val) = ⋃₀ C := by ext x; constructor · rintro ⟨s, ⟨y, rfl⟩, ⟨s, hs⟩, rfl⟩ refine ⟨_, hs, ?_⟩ @@ -1364,7 +1372,7 @@ theorem inj_on_iUnion_of_directed {s : ι → Set α} (hs : Directed (· ⊆ ·) theorem surjOn_sUnion {s : Set α} {T : Set (Set β)} {f : α → β} (H : ∀ t ∈ T, SurjOn f s t) : - SurjOn f s (⋃₀T) := fun _ ⟨t, ht, hx⟩ => H t ht hx + SurjOn f s (⋃₀ T) := fun _ ⟨t, ht, hx⟩ => H t ht hx theorem surjOn_iUnion {s : Set α} {t : ι → Set β} {f : α → β} (H : ∀ i, SurjOn f s (t i)) : SurjOn f s (⋃ i, t i) := @@ -1467,6 +1475,28 @@ theorem biUnion_image : ⋃ x ∈ f '' s, g x = ⋃ y ∈ s, g (f y) := theorem biInter_image : ⋂ x ∈ f '' s, g x = ⋂ y ∈ s, g (f y) := iInf_image +lemma biUnion_image2 (s : Set α) (t : Set β) (f : α → β → γ) (g : γ → Set δ) : + ⋃ c ∈ image2 f s t, g c = ⋃ a ∈ s, ⋃ b ∈ t, g (f a b) := iSup_image2 .. + +lemma biInter_image2 (s : Set α) (t : Set β) (f : α → β → γ) (g : γ → Set δ) : + ⋂ c ∈ image2 f s t, g c = ⋂ a ∈ s, ⋂ b ∈ t, g (f a b) := iInf_image2 .. + +lemma iUnion_inter_iUnion {ι κ : Sort*} (f : ι → Set α) (g : κ → Set α) : + (⋃ i, f i) ∩ ⋃ j, g j = ⋃ i, ⋃ j, f i ∩ g j := by simp_rw [iUnion_inter, inter_iUnion] + +lemma iInter_union_iInter {ι κ : Sort*} (f : ι → Set α) (g : κ → Set α) : + (⋂ i, f i) ∪ ⋂ j, g j = ⋂ i, ⋂ j, f i ∪ g j := by simp_rw [iInter_union, union_iInter] + +lemma iUnion₂_inter_iUnion₂ {ι₁ κ₁ : Sort*} {ι₂ : ι₁ → Sort*} {k₂ : κ₁ → Sort*} + (f : ∀ i₁, ι₂ i₁ → Set α) (g : ∀ j₁, k₂ j₁ → Set α) : + (⋃ i₁, ⋃ i₂, f i₁ i₂) ∩ ⋃ j₁, ⋃ j₂, g j₁ j₂ = ⋃ i₁, ⋃ i₂, ⋃ j₁, ⋃ j₂, f i₁ i₂ ∩ g j₁ j₂ := by + simp_rw [iUnion_inter, inter_iUnion] + +lemma iInter₂_union_iInter₂ {ι₁ κ₁ : Sort*} {ι₂ : ι₁ → Sort*} {k₂ : κ₁ → Sort*} + (f : ∀ i₁, ι₂ i₁ → Set α) (g : ∀ j₁, k₂ j₁ → Set α) : + (⋂ i₁, ⋂ i₂, f i₁ i₂) ∪ ⋂ j₁, ⋂ j₂, g j₁ j₂ = ⋂ i₁, ⋂ i₂, ⋂ j₁, ⋂ j₂, f i₁ i₂ ∪ g j₁ j₂ := by + simp_rw [iInter_union, union_iInter] + end Image section Preimage @@ -1492,7 +1522,7 @@ theorem image_sUnion {f : α → β} {s : Set (Set α)} : (f '' ⋃₀ s) = ⋃ exact ⟨a, ⟨t, ht₁, ht₂⟩, rfl⟩ @[simp] -theorem preimage_sUnion {f : α → β} {s : Set (Set β)} : f ⁻¹' ⋃₀s = ⋃ t ∈ s, f ⁻¹' t := by +theorem preimage_sUnion {f : α → β} {s : Set (Set β)} : f ⁻¹' ⋃₀ s = ⋃ t ∈ s, f ⁻¹' t := by rw [sUnion_eq_biUnion, preimage_iUnion₂] theorem preimage_iInter {f : α → β} {s : ι → Set β} : (f ⁻¹' ⋂ i, s i) = ⋂ i, f ⁻¹' s i := by @@ -1527,7 +1557,7 @@ theorem prod_iUnion {s : Set α} {t : ι → Set β} : (s ×ˢ ⋃ i, t i) = ⋃ theorem prod_iUnion₂ {s : Set α} {t : ∀ i, κ i → Set β} : (s ×ˢ ⋃ (i) (j), t i j) = ⋃ (i) (j), s ×ˢ t i j := by simp_rw [prod_iUnion] -theorem prod_sUnion {s : Set α} {C : Set (Set β)} : s ×ˢ ⋃₀C = ⋃₀((fun t => s ×ˢ t) '' C) := by +theorem prod_sUnion {s : Set α} {C : Set (Set β)} : s ×ˢ ⋃₀ C = ⋃₀ ((fun t => s ×ˢ t) '' C) := by simp_rw [sUnion_eq_biUnion, biUnion_image, prod_iUnion₂] theorem iUnion_prod_const {s : ι → Set α} {t : Set β} : (⋃ i, s i) ×ˢ t = ⋃ i, s i ×ˢ t := by @@ -1540,7 +1570,7 @@ theorem iUnion₂_prod_const {s : ∀ i, κ i → Set α} {t : Set β} : (⋃ (i) (j), s i j) ×ˢ t = ⋃ (i) (j), s i j ×ˢ t := by simp_rw [iUnion_prod_const] theorem sUnion_prod_const {C : Set (Set α)} {t : Set β} : - ⋃₀C ×ˢ t = ⋃₀((fun s : Set α => s ×ˢ t) '' C) := by + ⋃₀ C ×ˢ t = ⋃₀ ((fun s : Set α => s ×ˢ t) '' C) := by simp only [sUnion_eq_biUnion, iUnion₂_prod_const, biUnion_image] theorem iUnion_prod {ι ι' α β} (s : ι → Set α) (t : ι' → Set β) : @@ -1807,12 +1837,12 @@ theorem disjoint_iUnion₂_right {s : Set α} {t : ∀ i, κ i → Set α} : @[simp] theorem disjoint_sUnion_left {S : Set (Set α)} {t : Set α} : - Disjoint (⋃₀S) t ↔ ∀ s ∈ S, Disjoint s t := + Disjoint (⋃₀ S) t ↔ ∀ s ∈ S, Disjoint s t := sSup_disjoint_iff @[simp] theorem disjoint_sUnion_right {s : Set α} {S : Set (Set α)} : - Disjoint s (⋃₀S) ↔ ∀ t ∈ S, Disjoint s t := + Disjoint s (⋃₀ S) ↔ ∀ t ∈ S, Disjoint s t := disjoint_sSup_iff lemma biUnion_compl_eq_of_pairwise_disjoint_of_iUnion_eq_univ {ι : Type*} {Es : ι → Set α} From 8bd57d67caa56c16d165be48ea7309648270f309 Mon Sep 17 00:00:00 2001 From: smorel394 Date: Wed, 18 Dec 2024 16:44:13 +0000 Subject: [PATCH 770/829] feat(CategoryTheory/Shift/Opposite): commutation with shifts on the opposite functor (#20038) Given a `CommShift` structure on a functor `F` between categories endowed with shifts, construct the corresponding `CommShift` structure on `F.op` for the naive shifts on the opposite categories. (And vice versa.) Co-authored-by: smorel394 <67864981+smorel394@users.noreply.github.com> --- Mathlib/CategoryTheory/Shift/Opposite.lean | 65 +++++++++++++++++++++- 1 file changed, 64 insertions(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Shift/Opposite.lean b/Mathlib/CategoryTheory/Shift/Opposite.lean index f7f2818f38e80..fd57598711473 100644 --- a/Mathlib/CategoryTheory/Shift/Opposite.lean +++ b/Mathlib/CategoryTheory/Shift/Opposite.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Joël Riou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ -import Mathlib.CategoryTheory.Shift.Basic +import Mathlib.CategoryTheory.Shift.CommShift import Mathlib.CategoryTheory.Preadditive.Opposite /-! @@ -26,6 +26,8 @@ namespace CategoryTheory open Limits +section + variable (C : Type*) [Category C] (A : Type*) [AddMonoid A] [HasShift C A] namespace HasShift @@ -110,4 +112,65 @@ lemma oppositeShiftFunctorAdd'_hom_app : subst h simp only [shiftFunctorAdd'_eq_shiftFunctorAdd, oppositeShiftFunctorAdd_hom_app] +end + +variable {C D : Type*} [Category C] [Category D] (F : C ⥤ D) (A : Type*) [AddMonoid A] + [HasShift C A] [HasShift D A] + +namespace Functor + +/-- +Given a `CommShift` structure on `F`, this is the corresponding `CommShift` structure on +`F.op` (for the naive shifts on the opposite categories). +-/ +@[simps] +noncomputable def commShiftOp [CommShift F A] : + CommShift (C := OppositeShift C A) (D := OppositeShift D A) F.op A where + iso a := (NatIso.op (F.commShiftIso a)).symm + zero := by + simp only + rw [commShiftIso_zero] + ext + simp only [op_obj, comp_obj, Iso.symm_hom, NatIso.op_inv, NatTrans.op_app, + CommShift.isoZero_inv_app, op_comp, CommShift.isoZero_hom_app, op_map] + erw [oppositeShiftFunctorZero_inv_app, oppositeShiftFunctorZero_hom_app] + rfl + add a b := by + simp only + rw [commShiftIso_add] + ext + simp only [op_obj, comp_obj, Iso.symm_hom, NatIso.op_inv, NatTrans.op_app, + CommShift.isoAdd_inv_app, op_comp, Category.assoc, CommShift.isoAdd_hom_app, op_map] + erw [oppositeShiftFunctorAdd_inv_app, oppositeShiftFunctorAdd_hom_app] + rfl + +/-- +Given a `CommShift` structure on `F.op` (for the naive shifts on the opposite categories), +this is the corresponding `CommShift` structure on `F`. +-/ +@[simps] +noncomputable def commShiftUnop + [CommShift (C := OppositeShift C A) (D := OppositeShift D A) F.op A] : CommShift F A where + iso a := NatIso.removeOp (F.op.commShiftIso (C := OppositeShift C A) + (D := OppositeShift D A) a).symm + zero := by + simp only + rw [commShiftIso_zero] + ext + simp only [comp_obj, NatIso.removeOp_hom, Iso.symm_hom, NatTrans.removeOp_app, op_obj, + CommShift.isoZero_inv_app, op_map, unop_comp, Quiver.Hom.unop_op, CommShift.isoZero_hom_app] + erw [oppositeShiftFunctorZero_hom_app, oppositeShiftFunctorZero_inv_app] + rfl + add a b := by + simp only + rw [commShiftIso_add] + ext + simp only [comp_obj, NatIso.removeOp_hom, Iso.symm_hom, NatTrans.removeOp_app, op_obj, + CommShift.isoAdd_inv_app, op_map, unop_comp, Quiver.Hom.unop_op, Category.assoc, + CommShift.isoAdd_hom_app] + erw [oppositeShiftFunctorAdd_hom_app, oppositeShiftFunctorAdd_inv_app] + rfl + +end Functor + end CategoryTheory From 5138bb221a9833cd90740fcf41d4978be57ab0b3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 18 Dec 2024 17:34:56 +0000 Subject: [PATCH 771/829] refactor(Combinatorics/Additive): strengthen the classification of sets with no doubling (#20043) It turns out that the current statement is too weak for me to classify finite 1-approximate subgroups. From GrowthInGroups (LeanCamCombi) --- .../Additive/VerySmallDoubling.lean | 45 ++++++++++--------- 1 file changed, 24 insertions(+), 21 deletions(-) diff --git a/Mathlib/Combinatorics/Additive/VerySmallDoubling.lean b/Mathlib/Combinatorics/Additive/VerySmallDoubling.lean index b6015f656048d..eaa549d787651 100644 --- a/Mathlib/Combinatorics/Additive/VerySmallDoubling.lean +++ b/Mathlib/Combinatorics/Additive/VerySmallDoubling.lean @@ -25,38 +25,41 @@ open MulOpposite MulAction open scoped Pointwise RightActions namespace Finset -variable {G : Type*} [Group G] [DecidableEq G] {A : Finset G} +variable {G : Type*} [Group G] [DecidableEq G] {A : Finset G} {a : G} -/-- A set with no doubling is either empty or the translate of a subgroup. - -Precisely, if `A` has no doubling then there exists a subgroup `H` such `aH = Ha = A` for all -`a ∈ A`. -/ -@[to_additive "A set with no doubling is either empty or the translate of a subgroup. - -Precisely, if `A` has no doubling then there exists a subgroup `H` such `a + H = H + a = A` for all -`a ∈ A`."] -lemma exists_subgroup_of_no_doubling (hA : #(A * A) ≤ #A) : - ∃ H : Subgroup G, ∀ a ∈ A, a •> (H : Set G) = A ∧ (H : Set G) <• a = A := by +@[to_additive] +private lemma smul_stabilizer_of_no_doubling_aux (hA : #(A * A) ≤ #A) (ha : a ∈ A) : + a •> (stabilizer G A : Set G) = A ∧ (stabilizer G A : Set G) <• a = A := by have smul_A {a} (ha : a ∈ A) : a •> A = A * A := eq_of_subset_of_card_le (smul_finset_subset_mul ha) (by simpa) have A_smul {a} (ha : a ∈ A) : A <• a = A * A := eq_of_subset_of_card_le (op_smul_finset_subset_mul ha) (by simpa) - have smul_A_eq_op_smul_A {a} (ha : a ∈ A) : a •> A = A <• a := by rw [smul_A ha, A_smul ha] - have smul_A_eq_op_smul_A' {a} (ha : a ∈ A) : a⁻¹ •> A = A <• a⁻¹ := by - rw [inv_smul_eq_iff, smul_comm, smul_A_eq_op_smul_A ha, op_inv, inv_smul_smul] + have smul_A_eq_A_smul {a} (ha : a ∈ A) : a •> A = A <• a := by rw [smul_A ha, A_smul ha] + have mul_mem_A_comm {x a} (ha : a ∈ A) : x * a ∈ A ↔ a * x ∈ A := by + rw [← smul_mem_smul_finset_iff a, smul_A_eq_A_smul ha, ← op_smul_eq_mul, smul_comm, + smul_mem_smul_finset_iff, smul_eq_mul] let H := stabilizer G A have inv_smul_A {a} (ha : a ∈ A) : a⁻¹ • (A : Set G) = H := by ext x - refine ⟨?_, fun hx ↦ ?_⟩ - · rintro ⟨b, hb, rfl⟩ - simp [H, mul_smul, inv_smul_eq_iff, smul_A ha, smul_A hb] + rw [Set.mem_inv_smul_set_iff, smul_eq_mul] + refine ⟨fun hx ↦ ?_, fun hx ↦ ?_⟩ + · simpa [← smul_A ha, mul_smul] using smul_A hx · norm_cast - rwa [smul_A_eq_op_smul_A' ha, op_inv, mem_inv_smul_finset_iff, op_smul_eq_mul, ← smul_eq_mul, - ← mem_inv_smul_finset_iff, inv_mem hx] - refine ⟨H, fun a ha ↦ ⟨?_, ?_⟩⟩ + rwa [← mul_mem_A_comm ha, ← smul_eq_mul, ← mem_inv_smul_finset_iff, inv_mem hx] + refine ⟨?_, ?_⟩ · rw [← inv_smul_A ha, smul_inv_smul] · rw [← inv_smul_A ha, smul_comm] norm_cast - rw [← smul_A_eq_op_smul_A ha, inv_smul_smul] + rw [← smul_A_eq_A_smul ha, inv_smul_smul] + +/-- A non-empty set with no doubling is the left translate of its stabilizer. -/ +@[to_additive "A non-empty set with no doubling is the left-translate of its stabilizer."] +lemma smul_stabilizer_of_no_doubling (hA : #(A * A) ≤ #A) (ha : a ∈ A) : + a •> (stabilizer G A : Set G) = A := (smul_stabilizer_of_no_doubling_aux hA ha).1 + +/-- A non-empty set with no doubling is the right translate of its stabilizer. -/ +@[to_additive "A non-empty set with no doubling is the right translate of its stabilizer."] +lemma op_smul_stabilizer_of_no_doubling (hA : #(A * A) ≤ #A) (ha : a ∈ A) : + (stabilizer G A : Set G) <• a = A := (smul_stabilizer_of_no_doubling_aux hA ha).2 end Finset From ad5af7d032a4ddd1e7dea6bf1fe479b7d80a4d43 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Javier=20L=C3=B3pez-Contreras?= Date: Wed, 18 Dec 2024 17:43:34 +0000 Subject: [PATCH 772/829] feat(Complex): `reProdIm` lemmas (#19709) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From FLT Co-authored-by: Yaël Dillies Co-authored-by: Yaël Dillies --- Mathlib/Analysis/Complex/Basic.lean | 4 ++++ Mathlib/Analysis/Complex/ReImTopology.lean | 2 +- Mathlib/Data/Complex/Basic.lean | 17 +++++++++++++++-- 3 files changed, 20 insertions(+), 3 deletions(-) diff --git a/Mathlib/Analysis/Complex/Basic.lean b/Mathlib/Analysis/Complex/Basic.lean index 579ee76c75fee..05dc52915295f 100644 --- a/Mathlib/Analysis/Complex/Basic.lean +++ b/Mathlib/Analysis/Complex/Basic.lean @@ -716,4 +716,8 @@ lemma mem_slitPlane_of_norm_lt_one {z : ℂ} (hz : ‖z‖ < 1) : 1 + z ∈ slit end slitPlane +lemma _root_.IsCompact.reProdIm {s t : Set ℝ} (hs : IsCompact s) (ht : IsCompact t) : + IsCompact (s ×ℂ t) := + equivRealProdCLM.toHomeomorph.isCompact_preimage.2 (hs.prod ht) + end Complex diff --git a/Mathlib/Analysis/Complex/ReImTopology.lean b/Mathlib/Analysis/Complex/ReImTopology.lean index 1b2324b7f0db3..32c4888415dc3 100644 --- a/Mathlib/Analysis/Complex/ReImTopology.lean +++ b/Mathlib/Analysis/Complex/ReImTopology.lean @@ -150,7 +150,7 @@ theorem closure_reProdIm (s t : Set ℝ) : closure (s ×ℂ t) = closure s ×ℂ equivRealProdCLM.symm.toHomeomorph.preimage_closure] using @closure_prod_eq _ _ _ _ s t theorem interior_reProdIm (s t : Set ℝ) : interior (s ×ℂ t) = interior s ×ℂ interior t := by - rw [Set.reProdIm, Set.reProdIm, interior_inter, interior_preimage_re, interior_preimage_im] + rw [reProdIm, reProdIm, interior_inter, interior_preimage_re, interior_preimage_im] theorem frontier_reProdIm (s t : Set ℝ) : frontier (s ×ℂ t) = closure s ×ℂ frontier t ∪ frontier s ×ℂ closure t := by diff --git a/Mathlib/Data/Complex/Basic.lean b/Mathlib/Data/Complex/Basic.lean index c14f55bdce967..26342efc5f447 100644 --- a/Mathlib/Data/Complex/Basic.lean +++ b/Mathlib/Data/Complex/Basic.lean @@ -60,6 +60,9 @@ theorem ext : ∀ {z w : ℂ}, z.re = w.re → z.im = w.im → z = w attribute [local ext] Complex.ext +lemma «forall» {p : ℂ → Prop} : (∀ x, p x) ↔ ∀ a b, p ⟨a, b⟩ := by aesop +lemma «exists» {p : ℂ → Prop} : (∃ x, p x) ↔ ∃ a b, p ⟨a, b⟩ := by aesop + theorem re_surjective : Surjective re := fun x => ⟨⟨x, 0⟩, rfl⟩ theorem im_surjective : Surjective im := fun y => ⟨⟨0, y⟩, rfl⟩ @@ -106,11 +109,13 @@ instance canLift : CanLift ℂ ℝ (↑) fun z => z.im = 0 where /-- The product of a set on the real axis and a set on the imaginary axis of the complex plane, denoted by `s ×ℂ t`. -/ -def Set.reProdIm (s t : Set ℝ) : Set ℂ := +def reProdIm (s t : Set ℝ) : Set ℂ := re ⁻¹' s ∩ im ⁻¹' t +@[deprecated (since := "2024-12-03")] protected alias Set.reProdIm := reProdIm + @[inherit_doc] -infixl:72 " ×ℂ " => Set.reProdIm +infixl:72 " ×ℂ " => reProdIm theorem mem_reProdIm {z : ℂ} {s t : Set ℝ} : z ∈ s ×ℂ t ↔ z.re ∈ s ∧ z.im ∈ t := Iff.rfl @@ -855,6 +860,14 @@ lemma reProdIm_subset_iff' {s s₁ t t₁ : Set ℝ} : convert prod_subset_prod_iff exact reProdIm_subset_iff +variable {s t : Set ℝ} + +@[simp] lemma reProdIm_nonempty : (s ×ℂ t).Nonempty ↔ s.Nonempty ∧ t.Nonempty := by + simp [Set.Nonempty, reProdIm, Complex.exists] + +@[simp] lemma reProdIm_eq_empty : s ×ℂ t = ∅ ↔ s = ∅ ∨ t = ∅ := by + simp [← not_nonempty_iff_eq_empty, reProdIm_nonempty, -not_and, not_and_or] + end reProdIm open scoped Interval From b0466d1f33bcd06da8603dfe618c721422b37ac6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 18 Dec 2024 18:00:29 +0000 Subject: [PATCH 773/829] feat: boolean subalgebras (#20019) From GrowthInGroups (LeanCamCombi) --- Mathlib.lean | 1 + Mathlib/Order/BooleanSubalgebra.lean | 381 +++++++++++++++++++++++++++ Mathlib/Order/Sublattice.lean | 6 +- 3 files changed, 386 insertions(+), 2 deletions(-) create mode 100644 Mathlib/Order/BooleanSubalgebra.lean diff --git a/Mathlib.lean b/Mathlib.lean index c7b18fb3042f5..7e7f14bec1318 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3958,6 +3958,7 @@ import Mathlib.Order.Basic import Mathlib.Order.Birkhoff import Mathlib.Order.BooleanAlgebra import Mathlib.Order.BooleanGenerators +import Mathlib.Order.BooleanSubalgebra import Mathlib.Order.Booleanisation import Mathlib.Order.Bounded import Mathlib.Order.BoundedOrder.Basic diff --git a/Mathlib/Order/BooleanSubalgebra.lean b/Mathlib/Order/BooleanSubalgebra.lean new file mode 100644 index 0000000000000..ca4126280f883 --- /dev/null +++ b/Mathlib/Order/BooleanSubalgebra.lean @@ -0,0 +1,381 @@ +/- +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.Order.Sublattice + +/-! +# Boolean subalgebras + +This file defines boolean subalgebras. +-/ + +open Function Set + +variable {ι : Sort*} {α β γ : Type*} + +variable (α) in +/-- A boolean subalgebra of a boolean algebra is a set containing the bottom and top elements, and +closed under suprema, infima and complements. -/ +structure BooleanSubalgebra [BooleanAlgebra α] extends Sublattice α where + compl_mem' {a} : a ∈ carrier → aᶜ ∈ carrier + bot_mem' : ⊥ ∈ carrier + +namespace BooleanSubalgebra +section BooleanAlgebra +variable [BooleanAlgebra α] [BooleanAlgebra β] [BooleanAlgebra γ] {L M : BooleanSubalgebra α} + {f : BoundedLatticeHom α β} {s t : Set α} {a b : α} + +initialize_simps_projections BooleanSubalgebra (carrier → coe, as_prefix coe) + +instance instSetLike : SetLike (BooleanSubalgebra α) α where + coe L := L.carrier + coe_injective' L M h := by obtain ⟨⟨_, _⟩, _⟩ := L; congr + +lemma coe_inj : (L : Set α) = M ↔ L = M := SetLike.coe_set_eq + +@[simp] lemma supClosed (L : BooleanSubalgebra α) : SupClosed (L : Set α) := L.supClosed' +@[simp] lemma infClosed (L : BooleanSubalgebra α) : InfClosed (L : Set α) := L.infClosed' + +lemma compl_mem (ha : a ∈ L) : aᶜ ∈ L := L.compl_mem' ha +@[simp] lemma compl_mem_iff : aᶜ ∈ L ↔ a ∈ L := ⟨fun ha ↦ by simpa using compl_mem ha, compl_mem⟩ +@[simp] lemma bot_mem : ⊥ ∈ L := L.bot_mem' +@[simp] lemma top_mem : ⊤ ∈ L := by simpa using compl_mem L.bot_mem +lemma sup_mem (ha : a ∈ L) (hb : b ∈ L) : a ⊔ b ∈ L := L.supClosed ha hb +lemma inf_mem (ha : a ∈ L) (hb : b ∈ L) : a ⊓ b ∈ L := L.infClosed ha hb +lemma sdiff_mem (ha : a ∈ L) (hb : b ∈ L) : a \ b ∈ L := by + simpa [sdiff_eq] using L.infClosed ha (compl_mem hb) +lemma himp_mem (ha : a ∈ L) (hb : b ∈ L) : a ⇨ b ∈ L := by + simpa [himp_eq] using L.supClosed hb (compl_mem ha) + +lemma mem_carrier : a ∈ L.carrier ↔ a ∈ L := .rfl +@[simp] lemma mem_toSublattice : a ∈ L.toSublattice ↔ a ∈ L := .rfl +@[simp] lemma mem_mk {L : Sublattice α} (h_compl h_bot) : a ∈ mk L h_compl h_bot ↔ a ∈ L := .rfl +@[simp] lemma coe_mk (L : Sublattice α) (h_compl h_bot) : (mk L h_compl h_bot : Set α) = L := rfl +@[simp] lemma mk_le_mk {L M : Sublattice α} (hL_compl hL_bot hM_compl hM_bot) : + mk L hL_compl hL_bot ≤ mk M hM_compl hM_bot ↔ L ≤ M := .rfl +@[simp] lemma mk_lt_mk{L M : Sublattice α} (hL_compl hL_bot hM_compl hM_bot) : + mk L hL_compl hL_bot < mk M hM_compl hM_bot ↔ L < M := .rfl + +/-- Copy of a boolean subalgebra with a new `carrier` equal to the old one. Useful to fix +definitional equalities. -/ +protected def copy (L : BooleanSubalgebra α) (s : Set α) (hs : s = L) : BooleanSubalgebra α where + toSublattice := L.toSublattice.copy s <| by subst hs; rfl + compl_mem' := by subst hs; exact L.compl_mem' + bot_mem' := by subst hs; exact L.bot_mem' + +@[simp, norm_cast] +lemma coe_copy (L : BooleanSubalgebra α) (s : Set α) (hs) : L.copy s hs = s := rfl + +lemma copy_eq (L : BooleanSubalgebra α) (s : Set α) (hs) : L.copy s hs = L := + SetLike.coe_injective hs + +/-- Two boolean subalgebras are equal if they have the same elements. -/ +lemma ext : (∀ a, a ∈ L ↔ a ∈ M) → L = M := SetLike.ext + +/-- A boolean subalgebra of a lattice inherits a bottom element. -/ +instance instBotCoe : Bot L where bot := ⟨⊥, bot_mem⟩ + +/-- A boolean subalgebra of a lattice inherits a top element. -/ +instance instTopCoe : Top L where top := ⟨⊤, top_mem⟩ + +/-- A boolean subalgebra of a lattice inherits a supremum. -/ +instance instSupCoe : Max L where max a b := ⟨a ⊔ b, L.supClosed a.2 b.2⟩ + +/-- A boolean subalgebra of a lattice inherits an infimum. -/ +instance instInfCoe : Min L where min a b := ⟨a ⊓ b, L.infClosed a.2 b.2⟩ + +/-- A boolean subalgebra of a lattice inherits a complement. -/ +instance instHasComplCoe : HasCompl L where compl a := ⟨aᶜ, compl_mem a.2⟩ + +/-- A boolean subalgebra of a lattice inherits a difference. -/ +instance instSDiffCoe : SDiff L where sdiff a b := ⟨a \ b, sdiff_mem a.2 b.2⟩ + +/-- A boolean subalgebra of a lattice inherits a Heyting implication. -/ +instance instHImpCoe : HImp L where himp a b := ⟨a ⇨ b, himp_mem a.2 b.2⟩ + +@[simp, norm_cast] lemma val_bot : (⊥ : L) = (⊥ : α) := rfl +@[simp, norm_cast] lemma val_top : (⊤ : L) = (⊤ : α) := rfl +@[simp, norm_cast] lemma val_sup (a b : L) : a ⊔ b = (a : α) ⊔ b := rfl +@[simp, norm_cast] lemma val_inf (a b : L) : a ⊓ b = (a : α) ⊓ b := rfl +@[simp, norm_cast] lemma val_compl (a : L) : aᶜ = (a : α)ᶜ := rfl +@[simp, norm_cast] lemma val_sdiff (a b : L) : a \ b = (a : α) \ b := rfl +@[simp, norm_cast] lemma val_himp (a b : L) : a ⇨ b = (a : α) ⇨ b := rfl + +@[simp] lemma mk_bot : (⟨⊥, bot_mem⟩ : L) = ⊥ := rfl +@[simp] lemma mk_top : (⟨⊤, top_mem⟩ : L) = ⊤ := rfl +@[simp] lemma mk_sup_mk (a b : α) (ha hb) : (⟨a, ha⟩ ⊔ ⟨b, hb⟩ : L) = ⟨a ⊔ b, L.supClosed ha hb⟩ := + rfl +@[simp] lemma mk_inf_mk (a b : α) (ha hb) : (⟨a, ha⟩ ⊓ ⟨b, hb⟩ : L) = ⟨a ⊓ b, L.infClosed ha hb⟩ := + rfl +@[simp] lemma compl_mk (a : α) (ha) : (⟨a, ha⟩ : L)ᶜ = ⟨aᶜ, compl_mem ha⟩ := rfl +@[simp] lemma mk_sdiff_mk (a b : α) (ha hb) : (⟨a, ha⟩ \ ⟨b, hb⟩ : L) = ⟨a \ b, sdiff_mem ha hb⟩ := + rfl +@[simp] lemma mk_himp_mk (a b : α) (ha hb) : (⟨a, ha⟩ ⇨ ⟨b, hb⟩ : L) = ⟨a ⇨ b, himp_mem ha hb⟩ := + rfl + +/-- A boolean subalgebra of a lattice inherits a boolean algebra structure. -/ +instance instBooleanAlgebraCoe (L : BooleanSubalgebra α) : BooleanAlgebra L := + Subtype.coe_injective.booleanAlgebra _ val_sup val_inf val_top val_bot val_compl val_sdiff + val_himp + +/-- The natural lattice hom from a boolean subalgebra to the original lattice. -/ +def subtype (L : BooleanSubalgebra α) : BoundedLatticeHom L α where + toFun := ((↑) : L → α) + map_bot' := L.val_bot + map_top' := L.val_top + map_sup' := val_sup + map_inf' := val_inf + +@[simp, norm_cast] lemma coe_subtype (L : BooleanSubalgebra α) : L.subtype = ((↑) : L → α) := rfl +lemma subtype_apply (L : BooleanSubalgebra α) (a : L) : L.subtype a = a := rfl + +lemma subtype_injective (L : BooleanSubalgebra α) : Injective <| subtype L := Subtype.coe_injective + +/-- The inclusion homomorphism from a boolean subalgebra `L` to a bigger boolean subalgebra `M`. -/ +def inclusion (h : L ≤ M) : BoundedLatticeHom L M where + toFun := Set.inclusion h + map_bot' := rfl + map_top' := rfl + map_sup' _ _ := rfl + map_inf' _ _ := rfl + +@[simp] lemma coe_inclusion (h : L ≤ M) : inclusion h = Set.inclusion h := rfl +lemma inclusion_apply (h : L ≤ M) (a : L) : inclusion h a = Set.inclusion h a := rfl + +lemma inclusion_injective (h : L ≤ M) : Injective <| inclusion h := Set.inclusion_injective h + +@[simp] lemma inclusion_rfl (L : BooleanSubalgebra α) : inclusion le_rfl = .id L := rfl +@[simp] lemma subtype_comp_inclusion (h : L ≤ M) : M.subtype.comp (inclusion h) = L.subtype := rfl + +/-- The maximum boolean subalgebra of a lattice. -/ +instance instTop : Top (BooleanSubalgebra α) where + top.carrier := univ + top.bot_mem' := mem_univ _ + top.compl_mem' _ := mem_univ _ + top.supClosed' := supClosed_univ + top.infClosed' := infClosed_univ + +/-- The trivial boolean subalgebra of a lattice. -/ +instance instBot : Bot (BooleanSubalgebra α) where + bot.carrier := {⊥, ⊤} + bot.bot_mem' := by aesop + bot.compl_mem' := by aesop + bot.supClosed' _ := by aesop + bot.infClosed' _ := by aesop + +/-- The inf of two boolean subalgebras is their intersection. -/ +instance instInf : Min (BooleanSubalgebra α) where + min L M := { carrier := L ∩ M + bot_mem' := ⟨bot_mem, bot_mem⟩ + compl_mem' := fun ha ↦ ⟨compl_mem ha.1, compl_mem ha.2⟩ + supClosed' := L.supClosed.inter M.supClosed + infClosed' := L.infClosed.inter M.infClosed } + +/-- The inf of boolean subalgebras is their intersection. -/ +instance instInfSet : InfSet (BooleanSubalgebra α) where + sInf S := { carrier := ⋂ L ∈ S, L + bot_mem' := mem_iInter₂.2 fun _ _ ↦ bot_mem + compl_mem' := fun ha ↦ mem_iInter₂.2 fun L hL ↦ compl_mem <| mem_iInter₂.1 ha L hL + supClosed' := supClosed_sInter <| forall_mem_range.2 fun L ↦ supClosed_sInter <| + forall_mem_range.2 fun _ ↦ L.supClosed + infClosed' := infClosed_sInter <| forall_mem_range.2 fun L ↦ infClosed_sInter <| + forall_mem_range.2 fun _ ↦ L.infClosed } + +instance instInhabited : Inhabited (BooleanSubalgebra α) := ⟨⊥⟩ + +/-- The top boolean subalgebra is isomorphic to the original boolean algebra. + +This is the boolean subalgebra version of `Equiv.Set.univ α`. -/ +def topEquiv : (⊤ : BooleanSubalgebra α) ≃o α where + toEquiv := Equiv.Set.univ _ + map_rel_iff' := .rfl + +@[simp, norm_cast] lemma coe_top : (⊤ : BooleanSubalgebra α) = (univ : Set α) := rfl +@[simp, norm_cast] lemma coe_bot : (⊥ : BooleanSubalgebra α) = ({⊥, ⊤} : Set α) := rfl +@[simp, norm_cast] lemma coe_inf (L M : BooleanSubalgebra α) : L ⊓ M = (L : Set α) ∩ M := rfl + +@[simp, norm_cast] +lemma coe_sInf (S : Set (BooleanSubalgebra α)) : sInf S = ⋂ L ∈ S, (L : Set α) := rfl + +@[simp, norm_cast] +lemma coe_iInf (f : ι → BooleanSubalgebra α) : ⨅ i, f i = ⋂ i, (f i : Set α) := by simp [iInf] + +@[simp, norm_cast] lemma coe_eq_univ : L = (univ : Set α) ↔ L = ⊤ := by rw [← coe_top, coe_inj] + +@[simp] lemma mem_bot : a ∈ (⊥ : BooleanSubalgebra α) ↔ a = ⊥ ∨ a = ⊤ := .rfl +@[simp] lemma mem_top : a ∈ (⊤ : BooleanSubalgebra α) := mem_univ _ +@[simp] lemma mem_inf : a ∈ L ⊓ M ↔ a ∈ L ∧ a ∈ M := .rfl +@[simp] lemma mem_sInf {S : Set (BooleanSubalgebra α)} : a ∈ sInf S ↔ ∀ L ∈ S, a ∈ L := by + rw [← SetLike.mem_coe]; simp +@[simp] lemma mem_iInf {f : ι → BooleanSubalgebra α} : a ∈ ⨅ i, f i ↔ ∀ i, a ∈ f i := by + rw [← SetLike.mem_coe]; simp + +/-- BooleanSubalgebras of a lattice form a complete lattice. -/ +instance instCompleteLattice : CompleteLattice (BooleanSubalgebra α) where + bot := ⊥ + bot_le _S _a := by aesop + top := ⊤ + le_top _S a _ha := mem_top + inf := (· ⊓ ·) + le_inf _L _M _N hM hN _a ha := ⟨hM ha, hN ha⟩ + inf_le_left _L _M _a := And.left + inf_le_right _L _M _a := And.right + __ := completeLatticeOfInf (BooleanSubalgebra α) + fun _s ↦ IsGLB.of_image SetLike.coe_subset_coe isGLB_biInf + +instance [IsEmpty α] : Subsingleton (BooleanSubalgebra α) := SetLike.coe_injective.subsingleton +instance [IsEmpty α] : Unique (BooleanSubalgebra α) := uniqueOfSubsingleton ⊤ + +/-- The preimage of a boolean subalgebra along a bounded lattice homomorphism. -/ +def comap (f : BoundedLatticeHom α β) (L : BooleanSubalgebra β) : BooleanSubalgebra α where + carrier := f ⁻¹' L + bot_mem' := by simp + compl_mem' := by simp [map_compl'] + supClosed' := L.supClosed.preimage _ + infClosed' := L.infClosed.preimage _ + +@[simp, norm_cast] +lemma coe_comap (L : BooleanSubalgebra β) (f : BoundedLatticeHom α β) : L.comap f = f ⁻¹' L := rfl + +@[simp] lemma mem_comap {L : BooleanSubalgebra β} : a ∈ L.comap f ↔ f a ∈ L := .rfl + +lemma comap_mono : Monotone (comap f) := fun _ _ ↦ preimage_mono + +@[simp] lemma comap_id (L : BooleanSubalgebra α) : L.comap (BoundedLatticeHom.id _) = L := rfl + +@[simp] lemma comap_comap (L : BooleanSubalgebra γ) (g : BoundedLatticeHom β γ) + (f : BoundedLatticeHom α β) : (L.comap g).comap f = L.comap (g.comp f) := rfl + +/-- The image of a boolean subalgebra along a monoid homomorphism is a boolean subalgebra. -/ +def map (f : BoundedLatticeHom α β) (L : BooleanSubalgebra α) : BooleanSubalgebra β where + carrier := f '' L + bot_mem' := ⟨⊥, by simp⟩ + compl_mem' := by rintro _ ⟨a, ha, rfl⟩; exact ⟨aᶜ, by simpa [map_compl']⟩ + supClosed' := L.supClosed.image f + infClosed' := L.infClosed.image f + +@[simp] lemma coe_map (f : BoundedLatticeHom α β) (L : BooleanSubalgebra α) : + (L.map f : Set β) = f '' L := rfl + +@[simp] lemma mem_map {b : β} : b ∈ L.map f ↔ ∃ a ∈ L, f a = b := .rfl + +lemma mem_map_of_mem (f : BoundedLatticeHom α β) {a : α} : a ∈ L → f a ∈ L.map f := + mem_image_of_mem f + +lemma apply_coe_mem_map (f : BoundedLatticeHom α β) (a : L) : f a ∈ L.map f := + mem_map_of_mem f a.prop + +lemma map_mono : Monotone (map f) := fun _ _ ↦ image_subset _ + +@[simp] lemma map_id : L.map (.id α) = L := SetLike.coe_injective <| image_id _ + +@[simp] lemma map_map (g : BoundedLatticeHom β γ) (f : BoundedLatticeHom α β) : + (L.map f).map g = L.map (g.comp f) := SetLike.coe_injective <| image_image _ _ _ + +lemma mem_map_equiv {f : α ≃o β} {a : β} : a ∈ L.map f ↔ f.symm a ∈ L := Set.mem_image_equiv + +lemma apply_mem_map_iff (hf : Injective f) : f a ∈ L.map f ↔ a ∈ L := hf.mem_set_image + +lemma map_equiv_eq_comap_symm (f : α ≃o β) (L : BooleanSubalgebra α) : + L.map f = L.comap (f.symm : BoundedLatticeHom β α) := + SetLike.coe_injective <| f.toEquiv.image_eq_preimage L + +lemma comap_equiv_eq_map_symm (f : β ≃o α) (L : BooleanSubalgebra α) : + L.comap f = L.map (f.symm : BoundedLatticeHom α β) := (map_equiv_eq_comap_symm f.symm L).symm + +lemma map_symm_eq_iff_eq_map {M : BooleanSubalgebra β} {e : β ≃o α} : + L.map ↑e.symm = M ↔ L = M.map ↑e := by + simp_rw [← coe_inj]; exact (Equiv.eq_image_iff_symm_image_eq _ _ _).symm + +lemma map_le_iff_le_comap {f : BoundedLatticeHom α β} {M : BooleanSubalgebra β} : + L.map f ≤ M ↔ L ≤ M.comap f := image_subset_iff + +lemma gc_map_comap (f : BoundedLatticeHom α β) : GaloisConnection (map f) (comap f) := + fun _ _ ↦ map_le_iff_le_comap + +@[simp] lemma map_bot (f : BoundedLatticeHom α β) : (⊥ : BooleanSubalgebra α).map f = ⊥ := + (gc_map_comap f).l_bot + +lemma map_sup (f : BoundedLatticeHom α β) (L M : BooleanSubalgebra α) : + (L ⊔ M).map f = L.map f ⊔ M.map f := (gc_map_comap f).l_sup + +lemma map_iSup (f : BoundedLatticeHom α β) (L : ι → BooleanSubalgebra α) : + (⨆ i, L i).map f = ⨆ i, (L i).map f := (gc_map_comap f).l_iSup + +@[simp] lemma comap_top (f : BoundedLatticeHom α β) : (⊤ : BooleanSubalgebra β).comap f = ⊤ := + (gc_map_comap f).u_top + +lemma comap_inf (L M : BooleanSubalgebra β) (f : BoundedLatticeHom α β) : + (L ⊓ M).comap f = L.comap f ⊓ M.comap f := (gc_map_comap f).u_inf + +lemma comap_iInf (f : BoundedLatticeHom α β) (L : ι → BooleanSubalgebra β) : + (⨅ i, L i).comap f = ⨅ i, (L i).comap f := (gc_map_comap f).u_iInf + +lemma map_inf_le (L M : BooleanSubalgebra α) (f : BoundedLatticeHom α β) : + map f (L ⊓ M) ≤ map f L ⊓ map f M := map_mono.map_inf_le _ _ + +lemma le_comap_sup (L M : BooleanSubalgebra β) (f : BoundedLatticeHom α β) : + comap f L ⊔ comap f M ≤ comap f (L ⊔ M) := comap_mono.le_map_sup _ _ + +lemma le_comap_iSup (f : BoundedLatticeHom α β) (L : ι → BooleanSubalgebra β) : + ⨆ i, (L i).comap f ≤ (⨆ i, L i).comap f := comap_mono.le_map_iSup + +lemma map_inf (L M : BooleanSubalgebra α) (f : BoundedLatticeHom α β) (hf : Injective f) : + map f (L ⊓ M) = map f L ⊓ map f M := by + rw [← SetLike.coe_set_eq] + simp [Set.image_inter hf] + +lemma map_top (f : BoundedLatticeHom α β) (h : Surjective f) : BooleanSubalgebra.map f ⊤ = ⊤ := + SetLike.coe_injective <| by simp [h.range_eq] + +/-- The minimum boolean subalgebra containing a given set. -/ +def closure (s : Set α) : BooleanSubalgebra α := sInf {L | s ⊆ L} + +variable {s : Set α} + +lemma mem_closure {x : α} : x ∈ closure s ↔ ∀ ⦃L : BooleanSubalgebra α⦄, s ⊆ L → x ∈ L := mem_sInf + +@[aesop safe 20 apply (rule_sets := [SetLike])] +lemma subset_closure : s ⊆ closure s := fun _ hx ↦ mem_closure.2 fun _ hK ↦ hK hx + +@[simp] lemma closure_le : closure s ≤ L ↔ s ⊆ L := ⟨subset_closure.trans, fun h ↦ sInf_le h⟩ + +lemma closure_mono (hst : s ⊆ t) : closure s ≤ closure t := sInf_le_sInf fun _L ↦ hst.trans + +/-- An induction principle for closure membership. If `p` holds for `⊥` and all elements of `s`, and +is preserved under suprema and complement, then `p` holds for all elements of the closure of `s`. -/ +@[elab_as_elim] +lemma closure_bot_sup_induction {p : ∀ g ∈ closure s, Prop} (mem : ∀ x hx, p x (subset_closure hx)) + (bot : p ⊥ bot_mem) + (sup : ∀ x hx y hy, p x hx → p y hy → p (x ⊔ y) (supClosed _ hx hy)) + (compl : ∀ x hx, p x hx → p xᶜ (compl_mem hx)) {x} (hx : x ∈ closure s) : p x hx := + have inf ⦃x hx y hy⦄ (hx' : p x hx) (hy' : p y hy) : p (x ⊓ y) (infClosed _ hx hy) := by + simpa using compl _ _ <| sup _ _ _ _ (compl _ _ hx') (compl _ _ hy') + let L : BooleanSubalgebra α := + { carrier := { x | ∃ hx, p x hx } + supClosed' := fun _a ⟨_, ha⟩ _b ⟨_, hb⟩ ↦ ⟨_, sup _ _ _ _ ha hb⟩ + infClosed' := fun _a ⟨_, ha⟩ _b ⟨_, hb⟩ ↦ ⟨_, inf ha hb⟩ + bot_mem' := ⟨_, bot⟩ + compl_mem' := fun ⟨_, hb⟩ ↦ ⟨_, compl _ _ hb⟩ } + closure_le (L := L).mpr (fun y hy ↦ ⟨subset_closure hy, mem y hy⟩) hx |>.elim fun _ ↦ id + +end BooleanAlgebra + +section CompleteBooleanAlgebra +variable [CompleteBooleanAlgebra α] {L : BooleanSubalgebra α} {f : ι → α} {s : Set α} + +lemma iSup_mem [Finite ι] (hf : ∀ i, f i ∈ L) : ⨆ i, f i ∈ L := L.supClosed.iSup_mem bot_mem hf +lemma iInf_mem [Finite ι] (hf : ∀ i, f i ∈ L) : ⨅ i, f i ∈ L := L.infClosed.iInf_mem top_mem hf +lemma sSup_mem (hs : s.Finite) (hsL : s ⊆ L) : sSup s ∈ L := L.supClosed.sSup_mem hs bot_mem hsL +lemma sInf_mem (hs : s.Finite) (hsL : s ⊆ L) : sInf s ∈ L := L.infClosed.sInf_mem hs top_mem hsL + +lemma biSup_mem {ι : Type*} {t : Set ι} {f : ι → α} (ht : t.Finite) (hf : ∀ i ∈ t, f i ∈ L) : + ⨆ i ∈ t, f i ∈ L := L.supClosed.biSup_mem ht bot_mem hf + +lemma biInf_mem {ι : Type*} {t : Set ι} {f : ι → α} (ht : t.Finite) (hf : ∀ i ∈ t, f i ∈ L) : + ⨅ i ∈ t, f i ∈ L := L.infClosed.biInf_mem ht top_mem hf + +end CompleteBooleanAlgebra +end BooleanSubalgebra diff --git a/Mathlib/Order/Sublattice.lean b/Mathlib/Order/Sublattice.lean index 085366a96f497..2b6d406832661 100644 --- a/Mathlib/Order/Sublattice.lean +++ b/Mathlib/Order/Sublattice.lean @@ -34,7 +34,7 @@ structure Sublattice where variable {α β γ} namespace Sublattice -variable {L M : Sublattice α} {f : LatticeHom α β} {s t : Set α} {a : α} +variable {L M : Sublattice α} {f : LatticeHom α β} {s t : Set α} {a b : α} instance instSetLike : SetLike (Sublattice α) α where coe L := L.carrier @@ -52,6 +52,8 @@ lemma coe_inj : (L : Set α) = M ↔ L = M := SetLike.coe_set_eq @[simp] lemma supClosed (L : Sublattice α) : SupClosed (L : Set α) := L.supClosed' @[simp] lemma infClosed (L : Sublattice α) : InfClosed (L : Set α) := L.infClosed' +lemma sup_mem (ha : a ∈ L) (hb : b ∈ L) : a ⊔ b ∈ L := L.supClosed ha hb +lemma inf_mem (ha : a ∈ L) (hb : b ∈ L) : a ⊓ b ∈ L := L.infClosed ha hb @[simp] lemma isSublattice (L : Sublattice α) : IsSublattice (L : Set α) := ⟨L.supClosed, L.infClosed⟩ @@ -154,7 +156,7 @@ instance instInfSet : InfSet (Sublattice α) where instance instInhabited : Inhabited (Sublattice α) := ⟨⊥⟩ -/-- The top sublattice is isomorphic to the lattice. +/-- The top sublattice is isomorphic to the original lattice. This is the sublattice version of `Equiv.Set.univ α`. -/ def topEquiv : (⊤ : Sublattice α) ≃o α where From 54018a349a07aab14b2d23b5d3f37eb013859e21 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 18 Dec 2024 18:21:54 +0000 Subject: [PATCH 774/829] doc(Data/Multiset/Basic): write a proper module doc (#20044) --- Mathlib/Data/Multiset/Basic.lean | 18 ++++++++++++++++-- 1 file changed, 16 insertions(+), 2 deletions(-) diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index 473e920a37352..7eabbce0e4d7e 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -14,9 +14,23 @@ import Mathlib.Order.Hom.Basic /-! # Multisets -These are implemented as the quotient of a list by permutations. + +Multisets are finite sets with duplicates allowed. They are implemented here as the quotient of +lists by permutation. This gives them computational content. + ## Notation -We define the global infix notation `::ₘ` for `Multiset.cons`. + +* `0`: The empty multiset. +* `{a}`: The multiset containing a single occurrence of `a`. +* `a ::ₘ s`: The multiset containing one more occurrence of `a` than `s` does. +* `s + t`: The multiset for which the number of occurrences of each `a` is the sum of the + occurrences of `a` in `s` and `t`. +* `s - t`: The multiset for which the number of occurrences of each `a` is the difference of the + occurrences of `a` in `s` and `t`. +* `s ∪ t`: The multiset for which the number of occurrences of each `a` is the max of the + occurrences of `a` in `s` and `t`. +* `s ∩ t`: The multiset for which the number of occurrences of each `a` is the min of the + occurrences of `a` in `s` and `t`. -/ -- No bundled ordered algebra should be required From 75f292bf0d5f49d567c8fc6854c4e54cd7b3adfb Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 18 Dec 2024 19:47:38 +0000 Subject: [PATCH 775/829] feat: growth of a finset in the quotient and intersection with a subgroup (#19412) From GrowthInGroups (LeanCamCombi) --- Mathlib.lean | 1 + .../Geometry/Group/Growth/QuotientInter.lean | 76 +++++++++++++++++++ Mathlib/Geometry/Group/Growth/README.md | 10 +++ Mathlib/Geometry/Group/README.md | 10 +++ 4 files changed, 97 insertions(+) create mode 100644 Mathlib/Geometry/Group/Growth/QuotientInter.lean create mode 100644 Mathlib/Geometry/Group/Growth/README.md create mode 100644 Mathlib/Geometry/Group/README.md diff --git a/Mathlib.lean b/Mathlib.lean index 7e7f14bec1318..15e3f119d136e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3069,6 +3069,7 @@ import Mathlib.Geometry.Euclidean.Sphere.Power import Mathlib.Geometry.Euclidean.Sphere.Ptolemy import Mathlib.Geometry.Euclidean.Sphere.SecondInter import Mathlib.Geometry.Euclidean.Triangle +import Mathlib.Geometry.Group.Growth.QuotientInter import Mathlib.Geometry.Manifold.Algebra.LeftInvariantDerivation import Mathlib.Geometry.Manifold.Algebra.LieGroup import Mathlib.Geometry.Manifold.Algebra.Monoid diff --git a/Mathlib/Geometry/Group/Growth/QuotientInter.lean b/Mathlib/Geometry/Group/Growth/QuotientInter.lean new file mode 100644 index 0000000000000..0aef50f969b21 --- /dev/null +++ b/Mathlib/Geometry/Group/Growth/QuotientInter.lean @@ -0,0 +1,76 @@ +/- +Copyright (c) 2024 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import Mathlib.Algebra.Group.Pointwise.Finset.Basic +import Mathlib.Algebra.Order.BigOperators.Group.Finset +import Mathlib.GroupTheory.QuotientGroup.Defs + +/-! +# Growth in the quotient and intersection with a subgroup + +For a group `G` and a subgroup `H ≤ G`, this file upper and lower bounds the growth of a finset by +its growth in `H` and `G ⧸ H`. +-/ + +open Finset Function +open scoped Pointwise + +namespace Finset +variable {G : Type*} [Group G] [DecidableEq G] {H : Subgroup G} [DecidablePred (· ∈ H)] [H.Normal] + {A : Finset G} {m n : ℕ} + +@[to_additive] +lemma card_pow_quotient_mul_pow_inter_subgroup_le : + #((A ^ m).image <| QuotientGroup.mk' H) * #{x ∈ A ^ n | x ∈ H} ≤ #(A ^ (m + n)) := by + set π := QuotientGroup.mk' H + let φ := invFunOn π (A ^ m) + have hφ : Set.InjOn φ (π '' (A ^ m)) := invFunOn_injOn_image .. + have hφA {a} (ha : a ∈ π '' (A ^ m)) : φ a ∈ A ^ m := by + have := invFunOn_mem (by simpa using ha) + norm_cast at this + simpa using this + have hπφ {a} (ha : a ∈ π '' (A ^ m)) : π (φ a) = a := invFunOn_eq (by simpa using ha) + calc + #((A ^ m).image π) * #{x ∈ A ^ n | x ∈ H} + _ = #(((A ^ m).image π).image φ) * #{x ∈ A ^ n | x ∈ H} := by + rw [Finset.card_image_of_injOn (f := φ) (mod_cast hφ)] + _ ≤ #(((A ^ m).image π).image φ * {x ∈ A ^ n | x ∈ H}) := by + rw [Finset.card_mul_iff.2] + simp only [Set.InjOn, coe_image, coe_pow, coe_filter, Set.mem_prod, Set.mem_image, + exists_exists_and_eq_and, Set.mem_setOf_eq, and_imp, forall_exists_index, Prod.forall, + Prod.mk.injEq] + rintro _ a₁ b₁ hb₁ rfl - ha₁ _ a₂ b₂ hb₂ rfl - ha₂ hab + have hπa₁ : π a₁ = 1 := (QuotientGroup.eq_one_iff _).2 ha₁ + have hπa₂ : π a₂ = 1 := (QuotientGroup.eq_one_iff _).2 ha₂ + have hπb : π b₁ = π b₂ := by + simpa [hπφ, Set.mem_image_of_mem π, hb₁, hb₂, hπa₁, hπa₂] using congr(π $hab) + aesop + _ ≤ #(A ^ (m + n)) := by + gcongr + simp only [mul_subset_iff, mem_image, exists_exists_and_eq_and, Finset.mem_filter, and_imp, + forall_exists_index, forall_apply_eq_imp_iff₂, pow_add] + rintro a ha b hb - + exact mul_mem_mul (hφA <| Set.mem_image_of_mem _ <| mod_cast ha) hb + +@[to_additive] +lemma le_card_quotient_mul_sq_inter_subgroup (hAsymm : A⁻¹ = A) : + #A ≤ #(A.image <| QuotientGroup.mk' H) * #{x ∈ A ^ 2 | x ∈ H} := by + classical + set π := QuotientGroup.mk' H + rw [card_eq_sum_card_image π] + refine sum_le_card_nsmul _ _ _ <| forall_mem_image.2 fun a ha ↦ ?_ + calc + #{a' ∈ A | π a' = π a} + _ ≤ #({a' ∈ A | π a' = π a}⁻¹ * {a' ∈ A | π a' = π a}) := + card_le_card_mul_left ⟨a⁻¹, by simpa⟩ + _ ≤ #{x ∈ A⁻¹ * A | x ∈ H} := by + gcongr + simp only [mul_subset_iff, mem_inv', map_inv, mem_filter, and_imp] + rintro x hx hxa y hy hya + refine ⟨mul_mem_mul (by simpa) hy, (QuotientGroup.eq_one_iff _).1 (?_ : π _ = _)⟩ + simp [hya, ← hxa] + _ = #{x ∈ A ^ 2 | x ∈ H} := by simp [hAsymm, sq] + +end Finset diff --git a/Mathlib/Geometry/Group/Growth/README.md b/Mathlib/Geometry/Group/Growth/README.md new file mode 100644 index 0000000000000..6f07dca284151 --- /dev/null +++ b/Mathlib/Geometry/Group/Growth/README.md @@ -0,0 +1,10 @@ +# Growth of groups + +This subfolder is destined to contain results about the growth of groups. + +If $G$ is a finitely generated group, say $G = \langle S\rangle$ for some finite symmetric set $S$, then the growth of $G$ is defined to be the function $n \mapsto |S^n|$ where $S^n$ denotes pointwise multiplication of $n$ copies of $S$. Up to scaling, the growth does not depend on the symmetric generating set $S$. + +## Topics + +The results on growth of groups currently covered are: +* The growth of a group $G$ is roughly the growth of a normal subgroup $H \le G$ times the growth of $G / H$. diff --git a/Mathlib/Geometry/Group/README.md b/Mathlib/Geometry/Group/README.md new file mode 100644 index 0000000000000..19d79aff6e266 --- /dev/null +++ b/Mathlib/Geometry/Group/README.md @@ -0,0 +1,10 @@ +# Geometric group theory + +This subfolder is destined to contain the formalisation of geometric group theory. + +Geometric group theory is the study of finitely generated groups through their actions on geometric spaces. + +## Topics + +The geometric group theory topics currently covered are: +* Growth of groups From 933ad5493803f4c495f252678466ede6e1e2ee79 Mon Sep 17 00:00:00 2001 From: Bhavik Mehta Date: Wed, 18 Dec 2024 19:47:40 +0000 Subject: [PATCH 776/829] chore(Logic/ExistsUnique): amend usages of existsUnique to follow naming convention (#20028) As a definition `ExistsUnique`, the spelling of usages should be `existsUnique` rather than `exists_unique`. The majority of applications already conform to this (96), but there are some (30) exceptions. This PR fixes those. The large import is the import of Tactic.Alias to Logic.ExistsUnique, and is necessary if deprecations are to be performed here. --- Mathlib/Algebra/CharP/Defs.lean | 10 +++-- Mathlib/Algebra/Field/IsField.lean | 2 +- Mathlib/Algebra/Order/Antidiag/Nat.lean | 11 +++-- .../PrimeSpectrum/Basic.lean | 2 +- Mathlib/CategoryTheory/Extensive.lean | 4 +- .../Limits/Preserves/Ulift.lean | 10 +++-- .../CategoryTheory/Sites/CoverPreserving.lean | 2 +- Mathlib/CategoryTheory/Sites/CoversTop.lean | 12 +++--- Mathlib/CategoryTheory/Sites/IsSheafFor.lean | 2 +- Mathlib/CategoryTheory/Sites/Sheaf.lean | 13 +++--- Mathlib/CategoryTheory/Sites/SheafHom.lean | 2 +- Mathlib/Data/Fintype/Card.lean | 4 +- Mathlib/Data/Int/ModEq.lean | 10 +++-- Mathlib/Data/Set/Pairwise/Lattice.lean | 2 +- Mathlib/Data/Setoid/Partition.lean | 4 +- Mathlib/Dynamics/Newton.lean | 11 +++-- Mathlib/GroupTheory/GroupAction/Blocks.lean | 2 +- Mathlib/LinearAlgebra/JordanChevalley.lean | 2 +- Mathlib/LinearAlgebra/LinearPMap.lean | 2 +- Mathlib/Logic/ExistsUnique.lean | 43 +++++++++++++------ Mathlib/Logic/Unique.lean | 9 +++- .../CanonicalEmbedding/FundamentalCone.lean | 5 ++- Mathlib/NumberTheory/Padics/RingHoms.lean | 10 +++-- .../RingTheory/Localization/Away/Basic.lean | 2 +- .../RingTheory/Polynomial/HilbertPoly.lean | 6 ++- .../UniqueFactorizationDomain/Basic.lean | 13 ++++-- .../UniqueFactorizationDomain/Defs.lean | 2 +- .../Topology/Algebra/Module/LinearPMap.lean | 4 +- Mathlib/Topology/Sheaves/PUnit.lean | 2 +- 29 files changed, 131 insertions(+), 72 deletions(-) diff --git a/Mathlib/Algebra/CharP/Defs.lean b/Mathlib/Algebra/CharP/Defs.lean index 093c600a4a310..00eef1858e846 100644 --- a/Mathlib/Algebra/CharP/Defs.lean +++ b/Mathlib/Algebra/CharP/Defs.lean @@ -146,25 +146,27 @@ lemma «exists» : ∃ p, CharP R p := of_not_not (not_not_of_not_imp <| Nat.find_spec (not_forall.1 H)), zero_mul]⟩⟩⟩ -lemma exists_unique : ∃! p, CharP R p := +lemma existsUnique : ∃! p, CharP R p := let ⟨c, H⟩ := CharP.exists R ⟨c, H, fun _y H2 => CharP.eq R H2 H⟩ +@[deprecated (since := "2024-12-17")] alias exists_unique := existsUnique + end NonAssocSemiring end CharP /-- Noncomputable function that outputs the unique characteristic of a semiring. -/ -noncomputable def ringChar [NonAssocSemiring R] : ℕ := Classical.choose (CharP.exists_unique R) +noncomputable def ringChar [NonAssocSemiring R] : ℕ := Classical.choose (CharP.existsUnique R) namespace ringChar variable [NonAssocSemiring R] lemma spec : ∀ x : ℕ, (x : R) = 0 ↔ ringChar R ∣ x := by - letI : CharP R (ringChar R) := (Classical.choose_spec (CharP.exists_unique R)).1 + letI : CharP R (ringChar R) := (Classical.choose_spec (CharP.existsUnique R)).1 exact CharP.cast_eq_zero_iff R (ringChar R) lemma eq (p : ℕ) [C : CharP R p] : ringChar R = p := - ((Classical.choose_spec (CharP.exists_unique R)).2 p C).symm + ((Classical.choose_spec (CharP.existsUnique R)).2 p C).symm instance charP : CharP R (ringChar R) := ⟨spec R⟩ diff --git a/Mathlib/Algebra/Field/IsField.lean b/Mathlib/Algebra/Field/IsField.lean index e9b67d20db248..f598ad7186426 100644 --- a/Mathlib/Algebra/Field/IsField.lean +++ b/Mathlib/Algebra/Field/IsField.lean @@ -77,7 +77,7 @@ a lemma that there is a unique inverse could be useful. theorem uniq_inv_of_isField (R : Type u) [Ring R] (hf : IsField R) : ∀ x : R, x ≠ 0 → ∃! y : R, x * y = 1 := by intro x hx - apply exists_unique_of_exists_of_unique + apply existsUnique_of_exists_of_unique · exact hf.mul_inv_cancel hx · intro y z hxy hxz calc diff --git a/Mathlib/Algebra/Order/Antidiag/Nat.lean b/Mathlib/Algebra/Order/Antidiag/Nat.lean index 88a8cc741a284..5db426bce4ec6 100644 --- a/Mathlib/Algebra/Order/Antidiag/Nat.lean +++ b/Mathlib/Algebra/Order/Antidiag/Nat.lean @@ -156,7 +156,7 @@ lemma image_piFinTwoEquiv_finMulAntidiag {n : ℕ} : ext x simp [(piFinTwoEquiv <| fun _ => ℕ).symm.surjective.exists] -lemma finMulAntidiag_exists_unique_prime_dvd {d n p : ℕ} (hn : Squarefree n) +lemma finMulAntidiag_existsUnique_prime_dvd {d n p : ℕ} (hn : Squarefree n) (hp : p ∈ n.primeFactorsList) (f : Fin d → ℕ) (hf : f ∈ finMulAntidiag d n) : ∃! i, p ∣ f i := by rw [mem_finMulAntidiag] at hf @@ -172,6 +172,9 @@ lemma finMulAntidiag_exists_unique_prime_dvd {d n p : ℕ} (hn : Squarefree n) ← Finset.mul_prod_erase _ _ (mem_erase.mpr ⟨hij, mem_univ _⟩), ← mul_assoc] apply Nat.dvd_mul_right +@[deprecated (since := "2024-12-17")] +alias finMulAntidiag_exists_unique_prime_dvd := finMulAntidiag_existsUnique_prime_dvd + private def primeFactorsPiBij (d n : ℕ) : ∀ f ∈ (n.primeFactors.pi fun _ => (univ : Finset <| Fin d)), Fin d → ℕ := fun f _ i => ∏ p ∈ {p ∈ n.primeFactors.attach | f p.1 p.2 = i} , p @@ -212,10 +215,10 @@ private theorem primeFactorsPiBij_inj (d n : ℕ) private theorem primeFactorsPiBij_surj (d n : ℕ) (hn : Squarefree n) (t : Fin d → ℕ) (ht : t ∈ finMulAntidiag d n) : ∃ (g : _) (hg : g ∈ pi n.primeFactors fun _ => univ), Nat.primeFactorsPiBij d n g hg = t := by - have exists_unique := fun (p : ℕ) (hp : p ∈ n.primeFactors) => - (finMulAntidiag_exists_unique_prime_dvd hn + have existsUnique := fun (p : ℕ) (hp : p ∈ n.primeFactors) => + (finMulAntidiag_existsUnique_prime_dvd hn (mem_primeFactors_iff_mem_primeFactorsList.mp hp) t ht) - choose f hf hf_unique using exists_unique + choose f hf hf_unique using existsUnique refine ⟨f, ?_, ?_⟩ · simp only [mem_pi, mem_univ, forall_true_iff] funext i diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index a209d03a31d28..71357ae506ade 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -799,7 +799,7 @@ lemma basicOpen_injOn_isIdempotentElem : @[stacks 00EE] lemma existsUnique_idempotent_basicOpen_eq_of_isClopen {s : Set (PrimeSpectrum R)} (hs : IsClopen s) : ∃! e : R, IsIdempotentElem e ∧ s = basicOpen e := by - refine exists_unique_of_exists_of_unique ?_ ?_; swap + refine existsUnique_of_exists_of_unique ?_ ?_; swap · rintro x y ⟨hx, rfl⟩ ⟨hy, eq⟩ exact basicOpen_injOn_isIdempotentElem hx hy (SetLike.ext' eq) cases subsingleton_or_nontrivial R diff --git a/Mathlib/CategoryTheory/Extensive.lean b/Mathlib/CategoryTheory/Extensive.lean index dd21c2f557352..18e02517b1d56 100644 --- a/Mathlib/CategoryTheory/Extensive.lean +++ b/Mathlib/CategoryTheory/Extensive.lean @@ -220,7 +220,7 @@ instance types.finitaryExtensive : FinitaryExtensive (Type u) := by intro x cases' h : s.fst x with val val · simp only [Types.binaryCoproductCocone_pt, Functor.const_obj_obj, Sum.inl.injEq, - exists_unique_eq'] + existsUnique_eq'] · apply_fun f at h cases ((congr_fun s.condition x).symm.trans h).trans (congr_fun hαY val : _).symm delta ExistsUnique at this @@ -235,7 +235,7 @@ instance types.finitaryExtensive : FinitaryExtensive (Type u) := by · apply_fun f at h cases ((congr_fun s.condition x).symm.trans h).trans (congr_fun hαX val : _).symm · simp only [Types.binaryCoproductCocone_pt, Functor.const_obj_obj, Sum.inr.injEq, - exists_unique_eq'] + existsUnique_eq'] delta ExistsUnique at this choose l hl hl' using this exact ⟨l, (funext hl).symm, Types.isTerminalPunit.hom_ext _ _, diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean b/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean index eb1da8048a51a..d2998c6990db4 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean @@ -105,20 +105,22 @@ lemma descSet_inter_of_ne (x y : lc.pt) (hn : x ≠ y) : descSet hc {x} ∩ desc rw [mem_descSet_singleton] at hx hy exact hn (hx ▸ hy) -lemma exists_unique_mem_descSet (x : c.pt) : ∃! y : lc.pt, x ∈ descSet hc {y} := - exists_unique_of_exists_of_unique +lemma existsUnique_mem_descSet (x : c.pt) : ∃! y : lc.pt, x ∈ descSet hc {y} := + existsUnique_of_exists_of_unique (Set.mem_iUnion.mp <| Set.eq_univ_iff_forall.mp (iUnion_descSet_singleton hc lc) x) fun y₁ y₂ h₁ h₂ ↦ by_contra fun hn ↦ Set.eq_empty_iff_forall_not_mem.1 (descSet_inter_of_ne hc lc y₁ y₂ hn) x ⟨h₁, h₂⟩ +@[deprecated (since := "2024-12-17")] alias exists_unique_mem_descSet := existsUnique_mem_descSet + /-- Given a colimit cocone in `Type u` and an arbitrary cocone over the diagram lifted to `Type (max u v)`, produce a function from the cocone point of the colimit cocone to the cocone point of the other cocone, that witnesses the colimit cocone also being a colimit in the higher universe. -/ -noncomputable def descFun (x : c.pt) : lc.pt := (exists_unique_mem_descSet hc lc x).exists.choose +noncomputable def descFun (x : c.pt) : lc.pt := (existsUnique_mem_descSet hc lc x).exists.choose lemma descFun_apply_spec {x : c.pt} {y : lc.pt} : descFun hc lc x = y ↔ x ∈ descSet hc {y} := - have hu := exists_unique_mem_descSet hc lc x + have hu := existsUnique_mem_descSet hc lc x have hm := hu.exists.choose_spec ⟨fun he ↦ he ▸ hm, hu.unique hm⟩ diff --git a/Mathlib/CategoryTheory/Sites/CoverPreserving.lean b/Mathlib/CategoryTheory/Sites/CoverPreserving.lean index 2e0f39951144f..08f728dd0bf0e 100644 --- a/Mathlib/CategoryTheory/Sites/CoverPreserving.lean +++ b/Mathlib/CategoryTheory/Sites/CoverPreserving.lean @@ -164,7 +164,7 @@ This result is basically . lemma Functor.isContinuous_of_coverPreserving (hF₁ : CompatiblePreserving.{w} K F) (hF₂ : CoverPreserving J K F) : Functor.IsContinuous.{w} F J K where op_comp_isSheaf_of_types G X S hS x hx := by - apply exists_unique_of_exists_of_unique + apply existsUnique_of_exists_of_unique · have H := (isSheaf_iff_isSheaf_of_type _ _).1 G.2 _ (hF₂.cover_preserve hS) exact ⟨H.amalgamate (x.functorPushforward F) (hx.functorPushforward hF₁), fun V f hf => (H.isAmalgamation (hx.functorPushforward hF₁) (F.map f) _).trans diff --git a/Mathlib/CategoryTheory/Sites/CoversTop.lean b/Mathlib/CategoryTheory/Sites/CoversTop.lean index fafa6768c9f15..30872ff792de9 100644 --- a/Mathlib/CategoryTheory/Sites/CoversTop.lean +++ b/Mathlib/CategoryTheory/Sites/CoversTop.lean @@ -14,7 +14,7 @@ When there is a terminal object `X : C`, then `J.CoversTop Y` holds iff `Sieve.ofObjects Y X` is covering for `J`. We introduce a notion of compatible family of elements on objects `Y` -and obtain `Presheaf.FamilyOfElementsOnObjects.IsCompatible.exists_unique_section` +and obtain `Presheaf.FamilyOfElementsOnObjects.IsCompatible.existsUnique_section` which asserts that if a presheaf of types is a sheaf, then any compatible family of elements on objects `Y` which cover the final object extends as a section of this presheaf. @@ -121,10 +121,10 @@ lemma familyOfElements_isCompatible (hx : x.IsCompatible) (X : C) : variable {J} -lemma exists_unique_section (hx : x.IsCompatible) (hY : J.CoversTop Y) (hF : IsSheaf J F) : +lemma existsUnique_section (hx : x.IsCompatible) (hY : J.CoversTop Y) (hF : IsSheaf J F) : ∃! (s : F.sections), ∀ (i : I), s.1 (Opposite.op (Y i)) = x i := by have H := (isSheaf_iff_isSheaf_of_type _ _).1 hF - apply exists_unique_of_exists_of_unique + apply existsUnique_of_exists_of_unique · let s := fun (X : C) => (H _ (hY X)).amalgamate _ (hx.familyOfElements_isCompatible X) have hs : ∀ {X : C} (i : I) (f : X ⟶ Y i), s X = F.map f.op (x i) := fun {X} i f => by @@ -147,15 +147,17 @@ lemma exists_unique_section (hx : x.IsCompatible) (hY : J.CoversTop Y) (hF : IsS · intro y₁ y₂ hy₁ hy₂ exact hY.sections_ext ⟨F, hF⟩ (fun i => by rw [hy₁, hy₂]) +@[deprecated (since := "2024-12-17")] alias exists_unique_section := existsUnique_section + variable (hx : x.IsCompatible) (hY : J.CoversTop Y) (hF : IsSheaf J F) /-- The section of a sheaf of types which lifts a compatible family of elements indexed by objects which cover the terminal object. -/ -noncomputable def section_ : F.sections := (hx.exists_unique_section hY hF).choose +noncomputable def section_ : F.sections := (hx.existsUnique_section hY hF).choose @[simp] lemma section_apply (i : I) : (hx.section_ hY hF).1 (Opposite.op (Y i)) = x i := - (hx.exists_unique_section hY hF).choose_spec.1 i + (hx.existsUnique_section hY hF).choose_spec.1 i end IsCompatible diff --git a/Mathlib/CategoryTheory/Sites/IsSheafFor.lean b/Mathlib/CategoryTheory/Sites/IsSheafFor.lean index 45b5e94026382..0d7effb4b905f 100644 --- a/Mathlib/CategoryTheory/Sites/IsSheafFor.lean +++ b/Mathlib/CategoryTheory/Sites/IsSheafFor.lean @@ -533,7 +533,7 @@ theorem isSeparatedFor_and_exists_isAmalgamation_iff_isSheafFor : intro x constructor · intro z hx - exact exists_unique_of_exists_of_unique (z.2 hx) z.1 + exact existsUnique_of_exists_of_unique (z.2 hx) z.1 · intro h refine ⟨?_, ExistsUnique.exists ∘ h⟩ intro t₁ t₂ ht₁ ht₂ diff --git a/Mathlib/CategoryTheory/Sites/Sheaf.lean b/Mathlib/CategoryTheory/Sites/Sheaf.lean index 6900396ec52a3..a94ed5480644d 100644 --- a/Mathlib/CategoryTheory/Sites/Sheaf.lean +++ b/Mathlib/CategoryTheory/Sites/Sheaf.lean @@ -149,12 +149,12 @@ theorem isLimit_iff_isSheafFor : · intro hu E x hx specialize hu hx.cone rw [(homEquivAmalgamation hx).uniqueCongr.nonempty_congr] at hu - exact (unique_subtype_iff_exists_unique _).1 hu + exact (unique_subtype_iff_existsUnique _).1 hu · rintro h ⟨E, π⟩ let eqv := conesEquivSieveCompatibleFamily P S (op E) rw [← eqv.left_inv π] erw [(homEquivAmalgamation (eqv π).2).uniqueCongr.nonempty_congr] - rw [unique_subtype_iff_exists_unique] + rw [unique_subtype_iff_existsUnique] exact h _ _ (eqv π).2 /-- Given sieve `S` and presheaf `P : Cᵒᵖ ⥤ A`, their natural associated cone admits at most one @@ -267,21 +267,24 @@ variable {P : Cᵒᵖ ⥤ A} (hP : Presheaf.IsSheaf J P) {I : Type*} {S : C} {X a ≫ f i = b ≫ f j → x i ≫ P.map a.op = x j ≫ P.map b.op) include hP hf hx -lemma IsSheaf.exists_unique_amalgamation_ofArrows : +lemma IsSheaf.existsUnique_amalgamation_ofArrows : ∃! (g : E ⟶ P.obj (op S)), ∀ (i : I), g ≫ P.map (f i).op = x i := (Presieve.isSheafFor_arrows_iff _ _).1 ((Presieve.isSheafFor_iff_generate _).2 (hP E _ hf)) x (fun _ _ _ _ _ w => hx _ _ w) +@[deprecated (since := "2024-12-17")] +alias IsSheaf.exists_unique_amalgamation_ofArrows := IsSheaf.existsUnique_amalgamation_ofArrows + /-- If `P : Cᵒᵖ ⥤ A` is a sheaf and `f i : X i ⟶ S` is a covering family, then a morphism `E ⟶ P.obj (op S)` can be constructed from a compatible family of morphisms `x : E ⟶ P.obj (op (X i))`. -/ def IsSheaf.amalgamateOfArrows : E ⟶ P.obj (op S) := - (hP.exists_unique_amalgamation_ofArrows f hf x hx).choose + (hP.existsUnique_amalgamation_ofArrows f hf x hx).choose @[reassoc (attr := simp)] lemma IsSheaf.amalgamateOfArrows_map (i : I) : hP.amalgamateOfArrows f hf x hx ≫ P.map (f i).op = x i := - (hP.exists_unique_amalgamation_ofArrows f hf x hx).choose_spec.1 i + (hP.existsUnique_amalgamation_ofArrows f hf x hx).choose_spec.1 i end diff --git a/Mathlib/CategoryTheory/Sites/SheafHom.lean b/Mathlib/CategoryTheory/Sites/SheafHom.lean index d49a0a625c7ff..1b798e48eb161 100644 --- a/Mathlib/CategoryTheory/Sites/SheafHom.lean +++ b/Mathlib/CategoryTheory/Sites/SheafHom.lean @@ -165,7 +165,7 @@ open PresheafHom.IsSheafFor in lemma presheafHom_isSheafFor : Presieve.IsSheafFor (presheafHom F G) S.arrows := by intro x hx - apply exists_unique_of_exists_of_unique + apply existsUnique_of_exists_of_unique · refine ⟨ { app := fun Y => app hG x hx Y.unop.hom naturality := by diff --git a/Mathlib/Data/Fintype/Card.lean b/Mathlib/Data/Fintype/Card.lean index d9cf21d019a1f..0f1f730bb08da 100644 --- a/Mathlib/Data/Fintype/Card.lean +++ b/Mathlib/Data/Fintype/Card.lean @@ -540,13 +540,15 @@ theorem exists_pair_of_one_lt_card (h : 1 < card α) : ∃ a b : α, a ≠ b := theorem card_eq_one_of_forall_eq {i : α} (h : ∀ j, j = i) : card α = 1 := Fintype.card_eq_one_iff.2 ⟨i, h⟩ -theorem exists_unique_iff_card_one {α} [Fintype α] (p : α → Prop) [DecidablePred p] : +theorem existsUnique_iff_card_one {α} [Fintype α] (p : α → Prop) [DecidablePred p] : (∃! a : α, p a) ↔ #{x | p x} = 1 := by rw [Finset.card_eq_one] refine exists_congr fun x => ?_ simp only [forall_true_left, Subset.antisymm_iff, subset_singleton_iff', singleton_subset_iff, true_and, and_comm, mem_univ, mem_filter] +@[deprecated (since := "2024-12-17")] alias exists_unique_iff_card_one := existsUnique_iff_card_one + theorem one_lt_card [h : Nontrivial α] : 1 < Fintype.card α := Fintype.one_lt_card_iff_nontrivial.mpr h diff --git a/Mathlib/Data/Int/ModEq.lean b/Mathlib/Data/Int/ModEq.lean index 38b42a8cec85c..8374afec3980a 100644 --- a/Mathlib/Data/Int/ModEq.lean +++ b/Mathlib/Data/Int/ModEq.lean @@ -244,18 +244,22 @@ theorem mod_coprime {a b : ℕ} (hab : Nat.Coprime a b) : ∃ y : ℤ, a * y ≡ _ ≡ 1 [ZMOD ↑b] := by rw [← Nat.gcd_eq_gcd_ab, hgcd]; rfl ⟩ -theorem exists_unique_equiv (a : ℤ) {b : ℤ} (hb : 0 < b) : +theorem existsUnique_equiv (a : ℤ) {b : ℤ} (hb : 0 < b) : ∃ z : ℤ, 0 ≤ z ∧ z < b ∧ z ≡ a [ZMOD b] := ⟨a % b, emod_nonneg _ (ne_of_gt hb), by have : a % b < |b| := emod_lt _ (ne_of_gt hb) rwa [abs_of_pos hb] at this, by simp [ModEq]⟩ -theorem exists_unique_equiv_nat (a : ℤ) {b : ℤ} (hb : 0 < b) : ∃ z : ℕ, ↑z < b ∧ ↑z ≡ a [ZMOD b] := - let ⟨z, hz1, hz2, hz3⟩ := exists_unique_equiv a hb +@[deprecated (since := "2024-12-17")] alias exists_unique_equiv := existsUnique_equiv + +theorem existsUnique_equiv_nat (a : ℤ) {b : ℤ} (hb : 0 < b) : ∃ z : ℕ, ↑z < b ∧ ↑z ≡ a [ZMOD b] := + let ⟨z, hz1, hz2, hz3⟩ := existsUnique_equiv a hb ⟨z.natAbs, by constructor <;> rw [natAbs_of_nonneg hz1] <;> assumption⟩ +@[deprecated (since := "2024-12-17")] alias exists_unique_equiv_nat := existsUnique_equiv_nat + theorem mod_mul_right_mod (a b c : ℤ) : a % (b * c) % b = a % b := (mod_modEq _ _).of_mul_right _ diff --git a/Mathlib/Data/Set/Pairwise/Lattice.lean b/Mathlib/Data/Set/Pairwise/Lattice.lean index c71e6effc0488..644704f5ff134 100644 --- a/Mathlib/Data/Set/Pairwise/Lattice.lean +++ b/Mathlib/Data/Set/Pairwise/Lattice.lean @@ -154,7 +154,7 @@ theorem Pairwise.biUnion_injective (h₀ : Pairwise (Disjoint on f)) (h₁ : ∀ theorem pairwiseDisjoint_unique {y : α} (h_disjoint : PairwiseDisjoint s f) (hy : y ∈ (⋃ i ∈ s, f i)) : ∃! i, i ∈ s ∧ y ∈ f i := by - refine exists_unique_of_exists_of_unique ?ex ?unique + refine existsUnique_of_exists_of_unique ?ex ?unique · simpa only [mem_iUnion, exists_prop] using hy · rintro i j ⟨his, hi⟩ ⟨hjs, hj⟩ exact h_disjoint.elim his hjs <| not_disjoint_iff.mpr ⟨y, ⟨hi, hj⟩⟩ diff --git a/Mathlib/Data/Setoid/Partition.lean b/Mathlib/Data/Setoid/Partition.lean index ddc6a2e70e0d9..7781e1a597991 100644 --- a/Mathlib/Data/Setoid/Partition.lean +++ b/Mathlib/Data/Setoid/Partition.lean @@ -201,7 +201,7 @@ theorem IsPartition.pairwiseDisjoint {c : Set (Set α)} (hc : IsPartition c) : lemma _root_.Set.PairwiseDisjoint.isPartition_of_exists_of_ne_empty {α : Type*} {s : Set (Set α)} (h₁ : s.PairwiseDisjoint id) (h₂ : ∀ a : α, ∃ x ∈ s, a ∈ x) (h₃ : ∅ ∉ s) : Setoid.IsPartition s := by - refine ⟨h₃, fun a ↦ exists_unique_of_exists_of_unique (h₂ a) ?_⟩ + refine ⟨h₃, fun a ↦ existsUnique_of_exists_of_unique (h₂ a) ?_⟩ intro b₁ b₂ hb₁ hb₂ apply h₁.elim hb₁.1 hb₂.1 simp only [Set.not_disjoint_iff] @@ -212,7 +212,7 @@ theorem IsPartition.sUnion_eq_univ {c : Set (Set α)} (hc : IsPartition c) : ⋃ Set.mem_sUnion.2 <| let ⟨t, ht⟩ := hc.2 x ⟨t, by - simp only [exists_unique_iff_exists] at ht + simp only [existsUnique_iff_exists] at ht tauto⟩ /-- All elements of a partition of α are the equivalence class of some y ∈ α. -/ diff --git a/Mathlib/Dynamics/Newton.lean b/Mathlib/Dynamics/Newton.lean index 3a68a8bf036e7..8b6df8d843306 100644 --- a/Mathlib/Dynamics/Newton.lean +++ b/Mathlib/Dynamics/Newton.lean @@ -24,7 +24,7 @@ such as Hensel's lemma and Jordan-Chevalley decomposition. polynomial `P`. * `Polynomial.isFixedPt_newtonMap_of_isUnit_iff`: `x` is a fixed point for Newton iteration iff it is a root of `P` (provided `P'(x)` is a unit). - * `Polynomial.exists_unique_nilpotent_sub_and_aeval_eq_zero`: if `x` is almost a root of `P` in the + * `Polynomial.existsUnique_nilpotent_sub_and_aeval_eq_zero`: if `x` is almost a root of `P` in the sense that `P(x)` is nilpotent (and `P'(x)` is a unit) then we may write `x` as a sum `x = n + r` where `n` is nilpotent and `r` is a root of `P`. This can be used to prove the Jordan-Chevalley decomposition of linear endomorphims. @@ -77,7 +77,7 @@ theorem isFixedPt_newtonMap_of_isUnit_iff (h : IsUnit <| aeval x (derivative P)) rw [IsFixedPt, newtonMap_apply, sub_eq_self, Ring.inverse_mul_eq_iff_eq_mul _ _ _ h, mul_zero] /-- This is really an auxiliary result, en route to -`Polynomial.exists_unique_nilpotent_sub_and_aeval_eq_zero`. -/ +`Polynomial.existsUnique_nilpotent_sub_and_aeval_eq_zero`. -/ theorem aeval_pow_two_pow_dvd_aeval_iterate_newtonMap (h : IsNilpotent (aeval x P)) (h' : IsUnit (aeval x <| derivative P)) (n : ℕ) : (aeval x P) ^ (2 ^ n) ∣ aeval (P.newtonMap^[n] x) P := by @@ -103,11 +103,11 @@ unit) then we may write `x` as a sum `x = n + r` where `n` is nilpotent and `r` Moreover, `n` and `r` are unique. This can be used to prove the Jordan-Chevalley decomposition of linear endomorphims. -/ -theorem exists_unique_nilpotent_sub_and_aeval_eq_zero +theorem existsUnique_nilpotent_sub_and_aeval_eq_zero (h : IsNilpotent (aeval x P)) (h' : IsUnit (aeval x <| derivative P)) : ∃! r, IsNilpotent (x - r) ∧ aeval r P = 0 := by simp_rw [(neg_sub _ x).symm, isNilpotent_neg_iff] - refine exists_unique_of_exists_of_unique ?_ fun r₁ r₂ ⟨hr₁, hr₁'⟩ ⟨hr₂, hr₂'⟩ ↦ ?_ + refine existsUnique_of_exists_of_unique ?_ fun r₁ r₂ ⟨hr₁, hr₁'⟩ ⟨hr₂, hr₂'⟩ ↦ ?_ · -- Existence obtain ⟨n, hn⟩ := id h refine ⟨P.newtonMap^[n] x, isNilpotent_iterate_newtonMap_sub_of_isNilpotent h n, ?_⟩ @@ -125,4 +125,7 @@ theorem exists_unique_nilpotent_sub_and_aeval_eq_zero refine IsNilpotent.isUnit_add_left_of_commute ?_ this (Commute.all _ _) exact (Commute.all _ _).isNilpotent_mul_right <| (Commute.all _ _).isNilpotent_sub hr₂ hr₁ +@[deprecated (since := "2024-12-17")] +alias exists_unique_nilpotent_sub_and_aeval_eq_zero := existsUnique_nilpotent_sub_and_aeval_eq_zero + end Polynomial diff --git a/Mathlib/GroupTheory/GroupAction/Blocks.lean b/Mathlib/GroupTheory/GroupAction/Blocks.lean index 99a9a8640fb29..bb890e738c69d 100644 --- a/Mathlib/GroupTheory/GroupAction/Blocks.lean +++ b/Mathlib/GroupTheory/GroupAction/Blocks.lean @@ -345,7 +345,7 @@ theorem IsBlock.isBlockSystem [hGX : MulAction.IsPretransitive G X] obtain ⟨b : X, hb : b ∈ B⟩ := hBe obtain ⟨g, rfl⟩ := exists_smul_eq G b a use g • B - simp only [Set.smul_mem_smul_set_iff, hb, exists_unique_iff_exists, Set.mem_range, + simp only [Set.smul_mem_smul_set_iff, hb, existsUnique_iff_exists, Set.mem_range, exists_apply_eq_apply, exists_const, exists_prop, and_imp, forall_exists_index, forall_apply_eq_imp_iff, true_and] exact fun g' ha ↦ hB.smul_eq_smul_of_nonempty ⟨g • b, ha, ⟨b, hb, rfl⟩⟩ diff --git a/Mathlib/LinearAlgebra/JordanChevalley.lean b/Mathlib/LinearAlgebra/JordanChevalley.lean index b526d14e3ebcd..71814fa24fb3d 100644 --- a/Mathlib/LinearAlgebra/JordanChevalley.lean +++ b/Mathlib/LinearAlgebra/JordanChevalley.lean @@ -57,7 +57,7 @@ theorem exists_isNilpotent_isSemisimple_of_separable_of_dvd_pow {P : K[X]} {k : using (aeval f).congr_arg h refine isUnit_of_mul_eq_one_right (aeval ff b) _ (Subtype.ext_iff.mpr ?_) simpa [ff, coe_aeval_mk_apply] using h - obtain ⟨⟨s, mem⟩, ⟨⟨k, hk⟩, hss⟩, -⟩ := exists_unique_nilpotent_sub_and_aeval_eq_zero nil' sep' + obtain ⟨⟨s, mem⟩, ⟨⟨k, hk⟩, hss⟩, -⟩ := existsUnique_nilpotent_sub_and_aeval_eq_zero nil' sep' refine ⟨f - s, ?_, s, mem, ⟨k, ?_⟩, ?_, (sub_add_cancel f s).symm⟩ · exact sub_mem (self_mem_adjoin_singleton K f) mem · rw [Subtype.ext_iff] at hk diff --git a/Mathlib/LinearAlgebra/LinearPMap.lean b/Mathlib/LinearAlgebra/LinearPMap.lean index c0393fca46dd3..d460fe3990813 100644 --- a/Mathlib/LinearAlgebra/LinearPMap.lean +++ b/Mathlib/LinearAlgebra/LinearPMap.lean @@ -867,7 +867,7 @@ section SubmoduleToLinearPMap theorem existsUnique_from_graph {g : Submodule R (E × F)} (hg : ∀ {x : E × F} (_hx : x ∈ g) (_hx' : x.fst = 0), x.snd = 0) {a : E} (ha : a ∈ g.map (LinearMap.fst R E F)) : ∃! b : F, (a, b) ∈ g := by - refine exists_unique_of_exists_of_unique ?_ ?_ + refine existsUnique_of_exists_of_unique ?_ ?_ · convert ha simp intro y₁ y₂ hy₁ hy₂ diff --git a/Mathlib/Logic/ExistsUnique.lean b/Mathlib/Logic/ExistsUnique.lean index c41ef54d3dac5..51c99b37a5fd6 100644 --- a/Mathlib/Logic/ExistsUnique.lean +++ b/Mathlib/Logic/ExistsUnique.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Floris van Doorn -/ import Mathlib.Tactic.TypeStar +import Batteries.Tactic.Alias /-! # `ExistsUnique` @@ -78,10 +79,13 @@ theorem ExistsUnique.elim {p : α → Prop} {b : Prop} (h₂ : ∃! x, p x) (h₁ : ∀ x, p x → (∀ y, p y → y = x) → b) : b := Exists.elim h₂ (fun w hw ↦ h₁ w (And.left hw) (And.right hw)) -theorem exists_unique_of_exists_of_unique {p : α → Prop} +theorem existsUnique_of_exists_of_unique {p : α → Prop} (hex : ∃ x, p x) (hunique : ∀ y₁ y₂, p y₁ → p y₂ → y₁ = y₂) : ∃! x, p x := Exists.elim hex (fun x px ↦ ExistsUnique.intro x px (fun y (h : p y) ↦ hunique y x h px)) +@[deprecated (since := "2024-12-17")] +alias exists_unique_of_exists_of_unique := existsUnique_of_exists_of_unique + theorem ExistsUnique.exists {p : α → Prop} : (∃! x, p x) → ∃ x, p x | ⟨x, h, _⟩ => ⟨x, h⟩ theorem ExistsUnique.unique {p : α → Prop} @@ -96,37 +100,52 @@ theorem ExistsUnique.unique {p : α → Prop} theorem existsUnique_congr {p q : α → Prop} (h : ∀ a, p a ↔ q a) : (∃! a, p a) ↔ ∃! a, q a := exists_congr fun _ ↦ and_congr (h _) <| forall_congr' fun _ ↦ imp_congr_left (h _) -@[simp] theorem exists_unique_iff_exists [Subsingleton α] {p : α → Prop} : +@[simp] theorem existsUnique_iff_exists [Subsingleton α] {p : α → Prop} : (∃! x, p x) ↔ ∃ x, p x := ⟨fun h ↦ h.exists, Exists.imp fun x hx ↦ ⟨hx, fun y _ ↦ Subsingleton.elim y x⟩⟩ -theorem exists_unique_const {b : Prop} (α : Sort*) [i : Nonempty α] [Subsingleton α] : +@[deprecated (since := "2024-12-17")] alias exists_unique_iff_exists := existsUnique_iff_exists + +theorem existsUnique_const {b : Prop} (α : Sort*) [i : Nonempty α] [Subsingleton α] : (∃! _ : α, b) ↔ b := by simp -@[simp] theorem exists_unique_eq {a' : α} : ∃! a, a = a' := by +@[deprecated (since := "2024-12-17")] alias exists_unique_const := existsUnique_const + +@[simp] theorem existsUnique_eq {a' : α} : ∃! a, a = a' := by simp only [eq_comm, ExistsUnique, and_self, forall_eq', exists_eq'] -@[simp] theorem exists_unique_eq' {a' : α} : ∃! a, a' = a := by +@[deprecated (since := "2024-12-17")] alias exists_unique_eq := existsUnique_eq + +/-- The difference with `existsUnique_eq` is that the equality is reversed. -/ +@[simp] theorem existsUnique_eq' {a' : α} : ∃! a, a' = a := by simp only [ExistsUnique, and_self, forall_eq', exists_eq'] -theorem exists_unique_prop {p q : Prop} : (∃! _ : p, q) ↔ p ∧ q := by simp +@[deprecated (since := "2024-12-17")] alias exists_unique_eq' := existsUnique_eq' + +theorem existsUnique_prop {p q : Prop} : (∃! _ : p, q) ↔ p ∧ q := by simp + +@[deprecated (since := "2024-12-17")] alias exists_unique_prop := existsUnique_prop + +@[simp] theorem existsUnique_false : ¬∃! _ : α, False := fun ⟨_, h, _⟩ ↦ h + +@[deprecated (since := "2024-12-17")] alias exists_unique_false := existsUnique_false -@[simp] theorem exists_unique_false : ¬∃! _ : α, False := fun ⟨_, h, _⟩ ↦ h +theorem existsUnique_prop_of_true {p : Prop} {q : p → Prop} (h : p) : (∃! h' : p, q h') ↔ q h := + @existsUnique_const (q h) p ⟨h⟩ _ -theorem exists_unique_prop_of_true {p : Prop} {q : p → Prop} (h : p) : (∃! h' : p, q h') ↔ q h := - @exists_unique_const (q h) p ⟨h⟩ _ +@[deprecated (since := "2024-12-17")] alias exists_unique_prop_of_true := existsUnique_prop_of_true theorem ExistsUnique.elim₂ {p : α → Sort*} [∀ x, Subsingleton (p x)] {q : ∀ (x) (_ : p x), Prop} {b : Prop} (h₂ : ∃! x, ∃! h : p x, q x h) (h₁ : ∀ (x) (h : p x), q x h → (∀ (y) (hy : p y), q y hy → y = x) → b) : b := by - simp only [exists_unique_iff_exists] at h₂ + simp only [existsUnique_iff_exists] at h₂ apply h₂.elim exact fun x ⟨hxp, hxq⟩ H ↦ h₁ x hxp hxq fun y hyp hyq ↦ H y ⟨hyp, hyq⟩ theorem ExistsUnique.intro₂ {p : α → Sort*} [∀ x, Subsingleton (p x)] {q : ∀ (x : α) (_ : p x), Prop} (w : α) (hp : p w) (hq : q w hp) (H : ∀ (y) (hy : p y), q y hy → y = w) : ∃! x, ∃! hx : p x, q x hx := by - simp only [exists_unique_iff_exists] + simp only [existsUnique_iff_exists] exact ExistsUnique.intro w ⟨hp, hq⟩ fun y ⟨hyp, hyq⟩ ↦ H y hyp hyq theorem ExistsUnique.exists₂ {p : α → Sort*} {q : ∀ (x : α) (_ : p x), Prop} @@ -136,5 +155,5 @@ theorem ExistsUnique.exists₂ {p : α → Sort*} {q : ∀ (x : α) (_ : p x), P theorem ExistsUnique.unique₂ {p : α → Sort*} [∀ x, Subsingleton (p x)] {q : ∀ (x : α) (_ : p x), Prop} (h : ∃! x, ∃! hx : p x, q x hx) {y₁ y₂ : α} (hpy₁ : p y₁) (hqy₁ : q y₁ hpy₁) (hpy₂ : p y₂) (hqy₂ : q y₂ hpy₂) : y₁ = y₂ := by - simp only [exists_unique_iff_exists] at h + simp only [existsUnique_iff_exists] at h exact h.unique ⟨hpy₁, hqy₁⟩ ⟨hpy₂, hqy₂⟩ diff --git a/Mathlib/Logic/Unique.lean b/Mathlib/Logic/Unique.lean index 5c28b89ae00cf..adaadd587df02 100644 --- a/Mathlib/Logic/Unique.lean +++ b/Mathlib/Logic/Unique.lean @@ -54,17 +54,22 @@ attribute [class] Unique -- The simplifier can already prove this using `eq_iff_true_of_subsingleton` attribute [nolint simpNF] Unique.mk.injEq -theorem unique_iff_exists_unique (α : Sort u) : Nonempty (Unique α) ↔ ∃! _ : α, True := +theorem unique_iff_existsUnique (α : Sort u) : Nonempty (Unique α) ↔ ∃! _ : α, True := ⟨fun ⟨u⟩ ↦ ⟨u.default, trivial, fun a _ ↦ u.uniq a⟩, fun ⟨a, _, h⟩ ↦ ⟨⟨⟨a⟩, fun _ ↦ h _ trivial⟩⟩⟩ -theorem unique_subtype_iff_exists_unique {α} (p : α → Prop) : +@[deprecated (since := "2024-12-17")] alias unique_iff_exists_unique := unique_iff_existsUnique + +theorem unique_subtype_iff_existsUnique {α} (p : α → Prop) : Nonempty (Unique (Subtype p)) ↔ ∃! a, p a := ⟨fun ⟨u⟩ ↦ ⟨u.default.1, u.default.2, fun a h ↦ congr_arg Subtype.val (u.uniq ⟨a, h⟩)⟩, fun ⟨a, ha, he⟩ ↦ ⟨⟨⟨⟨a, ha⟩⟩, fun ⟨b, hb⟩ ↦ by congr exact he b hb⟩⟩⟩ +@[deprecated (since := "2024-12-17")] +alias unique_subtype_iff_exists_unique := unique_subtype_iff_existsUnique + /-- Given an explicit `a : α` with `Subsingleton α`, we can construct a `Unique α` instance. This is a def because the typeclass search cannot arbitrarily invent the `a : α` term. Nevertheless, these instances are all diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/FundamentalCone.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/FundamentalCone.lean index c305edd79b369..1863490ac9e87 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/FundamentalCone.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/FundamentalCone.lean @@ -256,12 +256,15 @@ theorem mem_integerSet {a : mixedSpace K} : /-- If `a` is in `integerSet`, then there is a *unique* algebraic integer in `𝓞 K` such that `mixedEmbedding K x = a`. -/ -theorem exists_unique_preimage_of_mem_integerSet {a : mixedSpace K} (ha : a ∈ integerSet K) : +theorem existsUnique_preimage_of_mem_integerSet {a : mixedSpace K} (ha : a ∈ integerSet K) : ∃! x : 𝓞 K, mixedEmbedding K x = a := by obtain ⟨_, ⟨x, rfl⟩⟩ := mem_integerSet.mp ha refine Function.Injective.existsUnique_of_mem_range ?_ (Set.mem_range_self x) exact (mixedEmbedding_injective K).comp RingOfIntegers.coe_injective +@[deprecated (since := "2024-12-17")] +alias exists_unique_preimage_of_mem_integerSet := existsUnique_preimage_of_mem_integerSet + theorem ne_zero_of_mem_integerSet (a : integerSet K) : (a : mixedSpace K) ≠ 0 := by by_contra! exact a.prop.1.2 (this.symm ▸ mixedEmbedding.norm.map_zero') diff --git a/Mathlib/NumberTheory/Padics/RingHoms.lean b/Mathlib/NumberTheory/Padics/RingHoms.lean index 8a7cb7bf0db3f..aabd86ef8f646 100644 --- a/Mathlib/NumberTheory/Padics/RingHoms.lean +++ b/Mathlib/NumberTheory/Padics/RingHoms.lean @@ -181,23 +181,25 @@ theorem exists_mem_range : ∃ n : ℕ, n < p ∧ x - n ∈ maximalIdeal ℤ_[p] apply max_lt hr simpa using hn -theorem exists_unique_mem_range : ∃! n : ℕ, n < p ∧ x - n ∈ maximalIdeal ℤ_[p] := by +theorem existsUnique_mem_range : ∃! n : ℕ, n < p ∧ x - n ∈ maximalIdeal ℤ_[p] := by obtain ⟨n, hn₁, hn₂⟩ := exists_mem_range x use n, ⟨hn₁, hn₂⟩, fun m ⟨hm₁, hm₂⟩ ↦ ?_ have := (zmod_congr_of_sub_mem_max_ideal x n m hn₂ hm₂).symm rwa [ZMod.natCast_eq_natCast_iff, ModEq, mod_eq_of_lt hn₁, mod_eq_of_lt hm₁] at this +@[deprecated (since := "2024-12-17")] alias exists_unique_mem_range := existsUnique_mem_range + /-- `zmod_repr x` is the unique natural number smaller than `p` satisfying `‖(x - zmod_repr x : ℤ_[p])‖ < 1`. -/ def zmodRepr : ℕ := - Classical.choose (exists_unique_mem_range x).exists + Classical.choose (existsUnique_mem_range x).exists theorem zmodRepr_spec : zmodRepr x < p ∧ x - zmodRepr x ∈ maximalIdeal ℤ_[p] := - Classical.choose_spec (exists_unique_mem_range x).exists + Classical.choose_spec (existsUnique_mem_range x).exists theorem zmodRepr_unique (y : ℕ) (hy₁ : y < p) (hy₂ : x - y ∈ maximalIdeal ℤ_[p]) : y = zmodRepr x := - have h := (Classical.choose_spec (exists_unique_mem_range x)).right + have h := (Classical.choose_spec (existsUnique_mem_range x)).right (h y ⟨hy₁, hy₂⟩).trans (h (zmodRepr x) (zmodRepr_spec x)).symm theorem zmodRepr_lt_p : zmodRepr x < p := diff --git a/Mathlib/RingTheory/Localization/Away/Basic.lean b/Mathlib/RingTheory/Localization/Away/Basic.lean index 259a7cba398d2..97d2e44e131c1 100644 --- a/Mathlib/RingTheory/Localization/Away/Basic.lean +++ b/Mathlib/RingTheory/Localization/Away/Basic.lean @@ -416,7 +416,7 @@ theorem existsUnique_algebraMap_eq_of_span_eq_top (s : Set R) (span_eq : Ideal.s h ⟨a, hts a.2⟩ ⟨b, hts b.2⟩) (Ideal.span_mono (fun _ ↦ .inr) mem) ⟨{a.1} ∪ t, by simp⟩ exact (congr_arg _ (uniq _ fun b ↦ eq ⟨b, .inr b.2⟩).symm).trans (eq ⟨a, .inl rfl⟩) have span_eq := (Ideal.eq_top_iff_one _).mpr mem - refine exists_unique_of_exists_of_unique ?_ fun x y hx hy ↦ + refine existsUnique_of_exists_of_unique ?_ fun x y hx hy ↦ algebraMap_injective_of_span_eq_top s span_eq (funext fun a ↦ (hx a).trans (hy a).symm) obtain ⟨s, rfl⟩ := finset_eq choose n r eq using fun a ↦ Away.surj a.1 (f a) diff --git a/Mathlib/RingTheory/Polynomial/HilbertPoly.lean b/Mathlib/RingTheory/Polynomial/HilbertPoly.lean index 66c5c57ec3060..32146950375d5 100644 --- a/Mathlib/RingTheory/Polynomial/HilbertPoly.lean +++ b/Mathlib/RingTheory/Polynomial/HilbertPoly.lean @@ -158,7 +158,7 @@ theorem coeff_mul_invOneSubPow_eq_hilbertPoly_eval /-- The polynomial satisfying the key property of `Polynomial.hilbertPoly p d` is unique. -/ -theorem exists_unique_hilbertPoly (p : F[X]) (d : ℕ) : +theorem existsUnique_hilbertPoly (p : F[X]) (d : ℕ) : ∃! h : F[X], ∃ N : ℕ, ∀ n > N, PowerSeries.coeff F n (p * invOneSubPow F d) = h.eval (n : F) := by use hilbertPoly p d; constructor @@ -171,6 +171,8 @@ theorem exists_unique_hilbertPoly (p : F[X]) (d : ℕ) : simp only [Set.mem_Ioi, sup_lt_iff, Set.mem_setOf_eq] at hn ⊢ rw [← coeff_mul_invOneSubPow_eq_hilbertPoly_eval d hn.2, hhN n hn.1] +@[deprecated (since := "2024-12-17")] alias exists_unique_hilbertPoly := existsUnique_hilbertPoly + /-- If `h : F[X]` and there exists some `N : ℕ` such that for any number `n : ℕ` bigger than `N` we have `PowerSeries.coeff F n (p * invOneSubPow F d) = h.eval (n : F)`, then `h` is exactly @@ -180,7 +182,7 @@ theorem eq_hilbertPoly_of_forall_coeff_eq_eval {p h : F[X]} {d : ℕ} (N : ℕ) (hhN : ∀ n > N, PowerSeries.coeff F n (p * invOneSubPow F d) = h.eval (n : F)) : h = hilbertPoly p d := - ExistsUnique.unique (exists_unique_hilbertPoly p d) ⟨N, hhN⟩ + ExistsUnique.unique (existsUnique_hilbertPoly p d) ⟨N, hhN⟩ ⟨p.natDegree, fun _ x => coeff_mul_invOneSubPow_eq_hilbertPoly_eval d x⟩ lemma hilbertPoly_mul_one_sub_succ (p : F[X]) (d : ℕ) : diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Basic.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Basic.lean index 31b9b8426dbed..9d04cd5c5c1f9 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain/Basic.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Basic.lean @@ -144,7 +144,7 @@ theorem prime_factors_irreducible [CancelCommMonoidWithZero α] {a : α} {f : Mu simp only [mul_one, Multiset.prod_cons, Multiset.prod_zero, hs0] at * exact ⟨Associated.symm ⟨u, hu⟩, rfl⟩ -theorem irreducible_iff_prime_of_exists_unique_irreducible_factors [CancelCommMonoidWithZero α] +theorem irreducible_iff_prime_of_existsUnique_irreducible_factors [CancelCommMonoidWithZero α] (eif : ∀ a : α, a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Irreducible b) ∧ f.prod ~ᵤ a) (uif : ∀ f g : Multiset α, @@ -183,6 +183,10 @@ theorem irreducible_iff_prime_of_exists_unique_irreducible_factors [CancelCommMo Or.inr <| hq.dvd_iff_dvd_left.2 <| hfb.2.dvd_iff_dvd_right.1 (Multiset.dvd_prod hqb)⟩, Prime.irreducible⟩ +@[deprecated (since := "2024-12-17")] +alias irreducible_iff_prime_of_exists_unique_irreducible_factors := + irreducible_iff_prime_of_existsUnique_irreducible_factors + namespace UniqueFactorizationMonoid variable [CancelCommMonoidWithZero α] @@ -401,7 +405,7 @@ end namespace UniqueFactorizationMonoid -theorem of_exists_unique_irreducible_factors [CancelCommMonoidWithZero α] +theorem of_existsUnique_irreducible_factors [CancelCommMonoidWithZero α] (eif : ∀ a : α, a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Irreducible b) ∧ f.prod ~ᵤ a) (uif : ∀ f g : Multiset α, @@ -411,7 +415,10 @@ theorem of_exists_unique_irreducible_factors [CancelCommMonoidWithZero α] UniqueFactorizationMonoid.of_exists_prime_factors (by convert eif using 7 - simp_rw [irreducible_iff_prime_of_exists_unique_irreducible_factors eif uif]) + simp_rw [irreducible_iff_prime_of_existsUnique_irreducible_factors eif uif]) + +@[deprecated (since := "2024-12-17")] +alias of_exists_unique_irreducible_factors := of_existsUnique_irreducible_factors variable {R : Type*} [CancelCommMonoidWithZero R] [UniqueFactorizationMonoid R] diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Defs.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Defs.lean index 28fef6a3998c7..9456594b84a19 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain/Defs.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Defs.lean @@ -120,7 +120,7 @@ Each element (except zero) is non-uniquely represented as a multiset of prime factors. To define a UFD using the definition in terms of multisets -of irreducible factors, use the definition `of_exists_unique_irreducible_factors` +of irreducible factors, use the definition `of_existsUnique_irreducible_factors` To define a UFD using the definition in terms of multisets of prime factors, use the definition `of_exists_prime_factors` diff --git a/Mathlib/Topology/Algebra/Module/LinearPMap.lean b/Mathlib/Topology/Algebra/Module/LinearPMap.lean index 1fccf24ee319b..7134c4a820669 100644 --- a/Mathlib/Topology/Algebra/Module/LinearPMap.lean +++ b/Mathlib/Topology/Algebra/Module/LinearPMap.lean @@ -28,7 +28,7 @@ underlying spaces are normed. * `LinearPMap.closable_iff_exists_closed_extension`: an unbounded operator is closable iff it has a closed extension. -* `LinearPMap.closable.exists_unique`: there exists a unique closure +* `LinearPMap.closable.existsUnique`: there exists a unique closure * `LinearPMap.closureHasCore`: the domain of `f` is a core of its closure ## References @@ -82,7 +82,7 @@ theorem IsClosable.leIsClosable {f g : E →ₗ.[R] F} (hf : f.IsClosable) (hfg /-- The closure is unique. -/ theorem IsClosable.existsUnique {f : E →ₗ.[R] F} (hf : f.IsClosable) : ∃! f' : E →ₗ.[R] F, f.graph.topologicalClosure = f'.graph := by - refine exists_unique_of_exists_of_unique hf fun _ _ hy₁ hy₂ => eq_of_eq_graph ?_ + refine existsUnique_of_exists_of_unique hf fun _ _ hy₁ hy₂ => eq_of_eq_graph ?_ rw [← hy₁, ← hy₂] open Classical in diff --git a/Mathlib/Topology/Sheaves/PUnit.lean b/Mathlib/Topology/Sheaves/PUnit.lean index 46e2456d0f3f3..671f5298e79c1 100644 --- a/Mathlib/Topology/Sheaves/PUnit.lean +++ b/Mathlib/Topology/Sheaves/PUnit.lean @@ -24,7 +24,7 @@ theorem isSheaf_of_isTerminal_of_indiscrete {X : TopCat.{w}} (hind : X.str = ⊤ (it : IsTerminal <| F.obj <| op ⊥) : F.IsSheaf := fun c U s hs => by obtain rfl | hne := eq_or_ne U ⊥ · intro _ _ - rw [@exists_unique_iff_exists _ ⟨fun _ _ => _⟩] + rw [@existsUnique_iff_exists _ ⟨fun _ _ => _⟩] · refine ⟨it.from _, fun U hU hs => IsTerminal.hom_ext ?_ _ _⟩ rwa [le_bot_iff.1 hU.le] · apply it.hom_ext From 054d75b44095e8390c2afd9744e30c3b2b6cbd42 Mon Sep 17 00:00:00 2001 From: Bryan Gin-ge Chen Date: Thu, 19 Dec 2024 08:51:56 +0000 Subject: [PATCH 777/829] fix: remove definitions from module doc that were moved to Defs.lean (#20048) --- Mathlib/GroupTheory/GroupAction/Basic.lean | 7 ------- 1 file changed, 7 deletions(-) diff --git a/Mathlib/GroupTheory/GroupAction/Basic.lean b/Mathlib/GroupTheory/GroupAction/Basic.lean index 7ea1e10af8050..347a3a65eba30 100644 --- a/Mathlib/GroupTheory/GroupAction/Basic.lean +++ b/Mathlib/GroupTheory/GroupAction/Basic.lean @@ -17,13 +17,6 @@ This file primarily concerns itself with orbits, stabilizers, and other objects actions. Despite this file being called `basic`, low-level helper lemmas for algebraic manipulation of `•` belong elsewhere. -## Main definitions - -* `MulAction.orbit` -* `MulAction.fixedPoints` -* `MulAction.fixedBy` -* `MulAction.stabilizer` - -/ From dac878dc2fde17aa9d7678bab2aed47c227a9063 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 19 Dec 2024 09:16:35 +0000 Subject: [PATCH 778/829] chore: move `transferTransversal` (#20042) The last section of `GroupTheory.Complement` is a mix of results purely about `ZMod` and results about the transfer transversal. The former can move to `Data.ZMod.Quotient`, while the latter can move to `GroupTheory.Transfer`. --- Mathlib/Data/ZMod/Quotient.lean | 25 +++++++- Mathlib/GroupTheory/Complement.lean | 81 +------------------------ Mathlib/GroupTheory/CosetCover.lean | 2 +- Mathlib/GroupTheory/HNNExtension.lean | 2 +- Mathlib/GroupTheory/Index.lean | 8 +++ Mathlib/GroupTheory/OrderOfElement.lean | 6 -- Mathlib/GroupTheory/Transfer.lean | 57 +++++++++++++++++ 7 files changed, 90 insertions(+), 91 deletions(-) diff --git a/Mathlib/Data/ZMod/Quotient.lean b/Mathlib/Data/ZMod/Quotient.lean index 1e52c60f93728..bcbc87a1d2aa0 100644 --- a/Mathlib/Data/ZMod/Quotient.lean +++ b/Mathlib/Data/ZMod/Quotient.lean @@ -3,9 +3,6 @@ Copyright (c) 2021 Anne Baanen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ -import Mathlib.Algebra.Group.Equiv.TypeTags -import Mathlib.Data.ZMod.Basic -import Mathlib.GroupTheory.GroupAction.Quotient import Mathlib.RingTheory.Ideal.Quotient.Operations import Mathlib.RingTheory.Int.Basic import Mathlib.RingTheory.ZMod @@ -208,3 +205,25 @@ lemma infinite_zpowers : (zpowers a : Set α).Infinite ↔ ¬IsOfFinOrder a := f protected alias ⟨_, IsOfFinOrder.finite_zpowers⟩ := finite_zpowers end Group + +namespace Subgroup +variable {G : Type*} [Group G] (H : Subgroup G) (g : G) + +open Equiv Function MulAction + +/-- Partition `G ⧸ H` into orbits of the action of `g : G`. -/ +noncomputable def quotientEquivSigmaZMod : + G ⧸ H ≃ Σq : orbitRel.Quotient (zpowers g) (G ⧸ H), ZMod (minimalPeriod (g • ·) q.out) := + (selfEquivSigmaOrbits (zpowers g) (G ⧸ H)).trans + (sigmaCongrRight fun q => orbitZPowersEquiv g q.out) + +lemma quotientEquivSigmaZMod_symm_apply (q : orbitRel.Quotient (zpowers g) (G ⧸ H)) + (k : ZMod (minimalPeriod (g • ·) q.out)) : + (quotientEquivSigmaZMod H g).symm ⟨q, k⟩ = g ^ (cast k : ℤ) • q.out := rfl + +lemma quotientEquivSigmaZMod_apply (q : orbitRel.Quotient (zpowers g) (G ⧸ H)) (k : ℤ) : + quotientEquivSigmaZMod H g (g ^ k • q.out) = ⟨q, k⟩ := by + rw [apply_eq_iff_eq_symm_apply, quotientEquivSigmaZMod_symm_apply, ZMod.coe_intCast, + zpow_smul_mod_minimalPeriod] + +end Subgroup diff --git a/Mathlib/GroupTheory/Complement.lean b/Mathlib/GroupTheory/Complement.lean index dc8d224d031ad..141b3bf2b8a77 100644 --- a/Mathlib/GroupTheory/Complement.lean +++ b/Mathlib/GroupTheory/Complement.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Thomas Browning. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning -/ -import Mathlib.Data.ZMod.Quotient +import Mathlib.GroupTheory.Index /-! # Complements @@ -655,82 +655,3 @@ theorem isComplement'_stabilizer {α : Type*} [MulAction G α] (a : α) rwa [Subtype.ext_iff, coe_one, coe_mul, ← self_eq_mul_left, mul_assoc (↑h) (↑h') g] at h1 end Subgroup - -namespace Subgroup - -open Equiv Function MemLeftTransversals MulAction MulAction.quotient ZMod - -universe u - -variable {G : Type u} [Group G] (H : Subgroup G) (g : G) - -/-- Partition `G ⧸ H` into orbits of the action of `g : G`. -/ -noncomputable def quotientEquivSigmaZMod : - G ⧸ H ≃ Σq : orbitRel.Quotient (zpowers g) (G ⧸ H), ZMod (minimalPeriod (g • ·) q.out) := - (selfEquivSigmaOrbits (zpowers g) (G ⧸ H)).trans - (sigmaCongrRight fun q => orbitZPowersEquiv g q.out) - -theorem quotientEquivSigmaZMod_symm_apply (q : orbitRel.Quotient (zpowers g) (G ⧸ H)) - (k : ZMod (minimalPeriod (g • ·) q.out)) : - (quotientEquivSigmaZMod H g).symm ⟨q, k⟩ = g ^ (cast k : ℤ) • q.out := - rfl - -theorem quotientEquivSigmaZMod_apply (q : orbitRel.Quotient (zpowers g) (G ⧸ H)) (k : ℤ) : - quotientEquivSigmaZMod H g (g ^ k • q.out) = ⟨q, k⟩ := by - rw [apply_eq_iff_eq_symm_apply, quotientEquivSigmaZMod_symm_apply, ZMod.coe_intCast, - zpow_smul_mod_minimalPeriod] - -/-- The transfer transversal as a function. Given a `⟨g⟩`-orbit `q₀, g • q₀, ..., g ^ (m - 1) • q₀` - in `G ⧸ H`, an element `g ^ k • q₀` is mapped to `g ^ k • g₀` for a fixed choice of - representative `g₀` of `q₀`. -/ -noncomputable def transferFunction : G ⧸ H → G := fun q => - g ^ (cast (quotientEquivSigmaZMod H g q).2 : ℤ) * (quotientEquivSigmaZMod H g q).1.out.out - -theorem transferFunction_apply (q : G ⧸ H) : - transferFunction H g q = - g ^ (cast (quotientEquivSigmaZMod H g q).2 : ℤ) * - (quotientEquivSigmaZMod H g q).1.out.out := - rfl - -theorem coe_transferFunction (q : G ⧸ H) : ↑(transferFunction H g q) = q := by - rw [transferFunction_apply, ← smul_eq_mul, Quotient.coe_smul_out, - ← quotientEquivSigmaZMod_symm_apply, Sigma.eta, symm_apply_apply] - -/-- The transfer transversal as a set. Contains elements of the form `g ^ k • g₀` for fixed choices - of representatives `g₀` of fixed choices of representatives `q₀` of `⟨g⟩`-orbits in `G ⧸ H`. -/ -def transferSet : Set G := - Set.range (transferFunction H g) - -theorem mem_transferSet (q : G ⧸ H) : transferFunction H g q ∈ transferSet H g := - ⟨q, rfl⟩ - -/-- The transfer transversal. Contains elements of the form `g ^ k • g₀` for fixed choices - of representatives `g₀` of fixed choices of representatives `q₀` of `⟨g⟩`-orbits in `G ⧸ H`. -/ -def transferTransversal : leftTransversals (H : Set G) := - ⟨transferSet H g, range_mem_leftTransversals (coe_transferFunction H g)⟩ - -theorem transferTransversal_apply (q : G ⧸ H) : - ↑(toEquiv (transferTransversal H g).2 q) = transferFunction H g q := - toEquiv_apply (coe_transferFunction H g) q - -theorem transferTransversal_apply' (q : orbitRel.Quotient (zpowers g) (G ⧸ H)) - (k : ZMod (minimalPeriod (g • ·) q.out)) : - ↑(toEquiv (transferTransversal H g).2 (g ^ (cast k : ℤ) • q.out)) = - g ^ (cast k : ℤ) * q.out.out := by - rw [transferTransversal_apply, transferFunction_apply, ← quotientEquivSigmaZMod_symm_apply, - apply_symm_apply] - -theorem transferTransversal_apply'' (q : orbitRel.Quotient (zpowers g) (G ⧸ H)) - (k : ZMod (minimalPeriod (g • ·) q.out)) : - ↑(toEquiv (g • transferTransversal H g).2 (g ^ (cast k : ℤ) • q.out)) = - if k = 0 then g ^ minimalPeriod (g • ·) q.out * q.out.out - else g ^ (cast k : ℤ) * q.out.out := by - rw [smul_apply_eq_smul_apply_inv_smul, transferTransversal_apply, transferFunction_apply, ← - mul_smul, ← zpow_neg_one, ← zpow_add, quotientEquivSigmaZMod_apply, smul_eq_mul, ← mul_assoc, - ← zpow_one_add, Int.cast_add, Int.cast_neg, Int.cast_one, intCast_cast, cast_id', id, ← - sub_eq_neg_add, cast_sub_one, add_sub_cancel] - by_cases hk : k = 0 - · rw [if_pos hk, if_pos hk, zpow_natCast] - · rw [if_neg hk, if_neg hk] - -end Subgroup diff --git a/Mathlib/GroupTheory/CosetCover.lean b/Mathlib/GroupTheory/CosetCover.lean index 745899e31384d..589e4d095bb97 100644 --- a/Mathlib/GroupTheory/CosetCover.lean +++ b/Mathlib/GroupTheory/CosetCover.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Antoine Chambert-Loir, Richard Copley. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Antoine Chambert-Loir, Richard Copley -/ - +import Mathlib.Algebra.Order.Ring.Rat import Mathlib.GroupTheory.Complement import Mathlib.LinearAlgebra.Basis.VectorSpace diff --git a/Mathlib/GroupTheory/HNNExtension.lean b/Mathlib/GroupTheory/HNNExtension.lean index 79dab5bd2e7c8..0134476792c78 100644 --- a/Mathlib/GroupTheory/HNNExtension.lean +++ b/Mathlib/GroupTheory/HNNExtension.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ - +import Mathlib.Algebra.Ring.Int.Units import Mathlib.GroupTheory.Coprod.Basic import Mathlib.GroupTheory.Complement diff --git a/Mathlib/GroupTheory/Index.lean b/Mathlib/GroupTheory/Index.lean index 95b2a5c3c1451..5bd39e3ccddbb 100644 --- a/Mathlib/GroupTheory/Index.lean +++ b/Mathlib/GroupTheory/Index.lean @@ -373,6 +373,14 @@ theorem relindex_eq_one : H.relindex K = 1 ↔ K ≤ H := theorem card_eq_one : Nat.card H = 1 ↔ H = ⊥ := H.relindex_bot_left ▸ relindex_eq_one.trans le_bot_iff +@[to_additive] +lemma inf_eq_bot_of_coprime (h : Nat.Coprime (Nat.card H) (Nat.card K)) : H ⊓ K = ⊥ := + card_eq_one.1 <| Nat.eq_one_of_dvd_coprimes h + (card_dvd_of_le inf_le_left) (card_dvd_of_le inf_le_right) + +@[deprecated (since := "2024-12-18")] +alias _root_.add_inf_eq_bot_of_coprime := AddSubgroup.inf_eq_bot_of_coprime + @[to_additive] theorem index_ne_zero_of_finite [hH : Finite (G ⧸ H)] : H.index ≠ 0 := by cases nonempty_fintype (G ⧸ H) diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index 492cc93e69891..a5f1e2aead8f6 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -969,12 +969,6 @@ lemma Nat.Coprime.pow_left_bijective {G} [Group G] (hn : (Nat.card G).Coprime n) Bijective (· ^ n : G → G) := (powCoprime hn).bijective -@[to_additive add_inf_eq_bot_of_coprime] -theorem inf_eq_bot_of_coprime {G : Type*} [Group G] {H K : Subgroup G} - (h : Nat.Coprime (Nat.card H) (Nat.card K)) : H ⊓ K = ⊥ := - card_eq_one.mp (Nat.eq_one_of_dvd_coprimes h - (card_dvd_of_le inf_le_left) (card_dvd_of_le inf_le_right)) - /- TODO: Generalise to `Submonoid.powers`. -/ @[to_additive] theorem image_range_orderOf [DecidableEq G] : diff --git a/Mathlib/GroupTheory/Transfer.lean b/Mathlib/GroupTheory/Transfer.lean index 8b40c9d992fa9..ed26a1a71afeb 100644 --- a/Mathlib/GroupTheory/Transfer.lean +++ b/Mathlib/GroupTheory/Transfer.lean @@ -73,6 +73,63 @@ theorem smul_diff_smul (g : G) : diff ϕ (g • S) (g • T) = diff ϕ S T := end leftTransversals +open Equiv Function MemLeftTransversals MulAction ZMod + +variable (g : G) + +variable (H) in +/-- The transfer transversal as a function. Given a `⟨g⟩`-orbit `q₀, g • q₀, ..., g ^ (m - 1) • q₀` + in `G ⧸ H`, an element `g ^ k • q₀` is mapped to `g ^ k • g₀` for a fixed choice of + representative `g₀` of `q₀`. -/ +noncomputable def transferFunction : G ⧸ H → G := fun q => + g ^ (cast (quotientEquivSigmaZMod H g q).2 : ℤ) * (quotientEquivSigmaZMod H g q).1.out.out + +lemma transferFunction_apply (q : G ⧸ H) : + transferFunction H g q = + g ^ (cast (quotientEquivSigmaZMod H g q).2 : ℤ) * + (quotientEquivSigmaZMod H g q).1.out.out := rfl + +lemma coe_transferFunction (q : G ⧸ H) : ↑(transferFunction H g q) = q := by + rw [transferFunction_apply, ← smul_eq_mul, Quotient.coe_smul_out, + ← quotientEquivSigmaZMod_symm_apply, Sigma.eta, symm_apply_apply] + +variable (H) in +/-- The transfer transversal as a set. Contains elements of the form `g ^ k • g₀` for fixed choices +of representatives `g₀` of fixed choices of representatives `q₀` of `⟨g⟩`-orbits in `G ⧸ H`. -/ +def transferSet : Set G := Set.range (transferFunction H g) + +lemma mem_transferSet (q : G ⧸ H) : transferFunction H g q ∈ transferSet H g := ⟨q, rfl⟩ + +variable (H) in +/-- The transfer transversal. Contains elements of the form `g ^ k • g₀` for fixed choices + of representatives `g₀` of fixed choices of representatives `q₀` of `⟨g⟩`-orbits in `G ⧸ H`. -/ +def transferTransversal : leftTransversals (H : Set G) := + ⟨transferSet H g, range_mem_leftTransversals (coe_transferFunction g)⟩ + +lemma transferTransversal_apply (q : G ⧸ H) : + ↑(toEquiv (transferTransversal H g).2 q) = transferFunction H g q := + toEquiv_apply (coe_transferFunction g) q + +lemma transferTransversal_apply' (q : orbitRel.Quotient (zpowers g) (G ⧸ H)) + (k : ZMod (minimalPeriod (g • ·) q.out)) : + ↑(toEquiv (transferTransversal H g).2 (g ^ (cast k : ℤ) • q.out)) = + g ^ (cast k : ℤ) * q.out.out := by + rw [transferTransversal_apply, transferFunction_apply, ← quotientEquivSigmaZMod_symm_apply, + apply_symm_apply] + +lemma transferTransversal_apply'' (q : orbitRel.Quotient (zpowers g) (G ⧸ H)) + (k : ZMod (minimalPeriod (g • ·) q.out)) : + ↑(toEquiv (g • transferTransversal H g).2 (g ^ (cast k : ℤ) • q.out)) = + if k = 0 then g ^ minimalPeriod (g • ·) q.out * q.out.out + else g ^ (cast k : ℤ) * q.out.out := by + rw [smul_apply_eq_smul_apply_inv_smul, transferTransversal_apply, transferFunction_apply, ← + mul_smul, ← zpow_neg_one, ← zpow_add, quotientEquivSigmaZMod_apply, smul_eq_mul, ← mul_assoc, + ← zpow_one_add, Int.cast_add, Int.cast_neg, Int.cast_one, intCast_cast, cast_id', id, ← + sub_eq_neg_add, cast_sub_one, add_sub_cancel] + by_cases hk : k = 0 + · rw [if_pos hk, if_pos hk, zpow_natCast] + · rw [if_neg hk, if_neg hk] + end Subgroup namespace MonoidHom From 3c80bbfb7c70c38203eae5a7cae5eae2c318cb20 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Martin=20Dvo=C5=99=C3=A1k?= Date: Thu, 19 Dec 2024 10:05:26 +0000 Subject: [PATCH 779/829] chore: unnamespace `Matrix.dotProduct` (#19910) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Moves: - `Matrix.dotProduct` -> `dotProduct` - `Matrix.dotProductᵣ` -> `dotProductᵣ` - The notation `⬝ᵥ` will no longer be scoped. - Lemmas about `dotProduct` that don't talk about matrices will be unnamespaced accordingly Discussion: https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/Matrix.2EdotProduct.20rename.3F --- Mathlib/Analysis/InnerProductSpace/PiL2.lean | 6 +- Mathlib/Data/Matrix/Mul.lean | 91 +++++++++++++++++-- Mathlib/Data/Matrix/Reflection.lean | 4 +- Mathlib/LinearAlgebra/AffineSpace/Matrix.lean | 2 +- Mathlib/LinearAlgebra/CrossProduct.lean | 4 +- .../LinearAlgebra/Matrix/BilinearForm.lean | 2 +- Mathlib/LinearAlgebra/Matrix/DotProduct.lean | 54 +++++++++-- Mathlib/LinearAlgebra/Matrix/Gershgorin.lean | 2 +- .../LinearAlgebra/Matrix/Nondegenerate.lean | 6 +- Mathlib/LinearAlgebra/Matrix/Orthogonal.lean | 4 +- Mathlib/LinearAlgebra/Matrix/PosDef.lean | 6 +- .../Matrix/SesquilinearForm.lean | 4 +- .../Matrix/SpecialLinearGroup.lean | 2 +- Mathlib/LinearAlgebra/Matrix/Spectrum.lean | 2 +- Mathlib/LinearAlgebra/Matrix/ToLin.lean | 2 +- .../Projectivization/Constructions.lean | 14 +-- Mathlib/Tactic/ScopedNS.lean | 4 +- scripts/nolints_prime_decls.txt | 6 +- 18 files changed, 162 insertions(+), 53 deletions(-) diff --git a/Mathlib/Analysis/InnerProductSpace/PiL2.lean b/Mathlib/Analysis/InnerProductSpace/PiL2.lean index 15e7d1d6c28ab..2f76e57c56c2e 100644 --- a/Mathlib/Analysis/InnerProductSpace/PiL2.lean +++ b/Mathlib/Analysis/InnerProductSpace/PiL2.lean @@ -185,11 +185,11 @@ theorem finrank_euclideanSpace_fin {n : ℕ} : 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) := + ⟪x, y⟫ = dotProduct (star <| WithLp.equiv _ _ x) (WithLp.equiv _ _ y) := rfl theorem EuclideanSpace.inner_piLp_equiv_symm (x y : ι → 𝕜) : - ⟪(WithLp.equiv 2 _).symm x, (WithLp.equiv 2 _).symm y⟫ = Matrix.dotProduct (star x) y := + ⟪(WithLp.equiv 2 _).symm x, (WithLp.equiv 2 _).symm y⟫ = dotProduct (star x) y := rfl /-- A finite, mutually orthogonal family of subspaces of `E`, which span `E`, induce an isometry @@ -1041,7 +1041,7 @@ local notation "⟪" x ", " y "⟫ₑ" => /-- The inner product of a row of `A` and a row of `B` is an entry of `B * Aᴴ`. -/ theorem inner_matrix_row_row [Fintype n] (A B : Matrix m n 𝕜) (i j : m) : ⟪A i, B j⟫ₑ = (B * Aᴴ) j i := by - simp_rw [EuclideanSpace.inner_piLp_equiv_symm, Matrix.mul_apply', Matrix.dotProduct_comm, + simp_rw [EuclideanSpace.inner_piLp_equiv_symm, Matrix.mul_apply', dotProduct_comm, Matrix.conjTranspose_apply, Pi.star_def] /-- The inner product of a column of `A` and a column of `B` is an entry of `Aᴴ * B`. -/ diff --git a/Mathlib/Data/Matrix/Mul.lean b/Mathlib/Data/Matrix/Mul.lean index 16985ffe09158..96502a2383140 100644 --- a/Mathlib/Data/Matrix/Mul.lean +++ b/Mathlib/Data/Matrix/Mul.lean @@ -11,10 +11,10 @@ import Mathlib.Data.Matrix.Diagonal /-! # Matrix multiplication -This file defines matrix multiplication +This file defines vector and matrix multiplication ## Main definitions - * `Matrix.dotProduct`: the dot product between two vectors + * `dotProduct`: the dot product between two vectors * `Matrix.mul`: multiplication of two matrices * `Matrix.mulVec`: multiplication of a matrix with a vector * `Matrix.vecMul`: multiplication of a vector with a matrix @@ -25,7 +25,7 @@ This file defines matrix multiplication The locale `Matrix` gives the following notation: -* `⬝ᵥ` for `Matrix.dotProduct` +* `⬝ᵥ` for `dotProduct` * `*ᵥ` for `Matrix.mulVec` * `ᵥ*` for `Matrix.vecMul` @@ -56,8 +56,6 @@ variable {R : Type*} {S : Type*} {α : Type v} {β : Type w} {γ : Type*} open Matrix -namespace Matrix - section DotProduct variable [Fintype m] [Fintype n] @@ -66,30 +64,42 @@ variable [Fintype m] [Fintype n] def dotProduct [Mul α] [AddCommMonoid α] (v w : m → α) : α := ∑ i, v i * w i +@[deprecated (since := "2024-12-12")] protected alias Matrix.dotProduct := dotProduct + /- The precedence of 72 comes immediately after ` • ` for `SMul.smul`, so that `r₁ • a ⬝ᵥ r₂ • b` is parsed as `(r₁ • a) ⬝ᵥ (r₂ • b)` here. -/ @[inherit_doc] -scoped infixl:72 " ⬝ᵥ " => Matrix.dotProduct +infixl:72 " ⬝ᵥ " => dotProduct theorem dotProduct_assoc [NonUnitalSemiring α] (u : m → α) (w : n → α) (v : Matrix m n α) : (fun j => u ⬝ᵥ fun i => v i j) ⬝ᵥ w = u ⬝ᵥ fun i => v i ⬝ᵥ w := by simpa [dotProduct, Finset.mul_sum, Finset.sum_mul, mul_assoc] using Finset.sum_comm +@[deprecated (since := "2024-12-12")] protected alias Matrix.dotProduct_assoc := dotProduct_assoc + theorem dotProduct_comm [AddCommMonoid α] [CommSemigroup α] (v w : m → α) : v ⬝ᵥ w = w ⬝ᵥ v := by simp_rw [dotProduct, mul_comm] +@[deprecated (since := "2024-12-12")] protected alias Matrix.dotProduct_comm := dotProduct_comm + @[simp] theorem dotProduct_pUnit [AddCommMonoid α] [Mul α] (v w : PUnit → α) : v ⬝ᵥ w = v ⟨⟩ * w ⟨⟩ := by simp [dotProduct] +@[deprecated (since := "2024-12-12")] protected alias Matrix.dotProduct_pUnit := dotProduct_pUnit + section MulOneClass variable [MulOneClass α] [AddCommMonoid α] theorem dotProduct_one (v : n → α) : v ⬝ᵥ 1 = ∑ i, v i := by simp [(· ⬝ᵥ ·)] +@[deprecated (since := "2024-12-12")] protected alias Matrix.dotProduct_one := dotProduct_one + theorem one_dotProduct (v : n → α) : 1 ⬝ᵥ v = ∑ i, v i := by simp [(· ⬝ᵥ ·)] +@[deprecated (since := "2024-12-12")] protected alias Matrix.one_dotProduct := one_dotProduct + end MulOneClass section NonUnitalNonAssocSemiring @@ -99,46 +109,70 @@ variable [NonUnitalNonAssocSemiring α] (u v w : m → α) (x y : n → α) @[simp] theorem dotProduct_zero : v ⬝ᵥ 0 = 0 := by simp [dotProduct] +@[deprecated (since := "2024-12-12")] protected alias Matrix.dotProduct_zero := dotProduct_zero + @[simp] theorem dotProduct_zero' : (v ⬝ᵥ fun _ => 0) = 0 := dotProduct_zero v +@[deprecated (since := "2024-12-12")] protected alias Matrix.dotProduct_zero' := dotProduct_zero' + @[simp] theorem zero_dotProduct : 0 ⬝ᵥ v = 0 := by simp [dotProduct] +@[deprecated (since := "2024-12-12")] protected alias Matrix.zero_dotProduct := zero_dotProduct + @[simp] theorem zero_dotProduct' : (fun _ => (0 : α)) ⬝ᵥ v = 0 := zero_dotProduct v +@[deprecated (since := "2024-12-12")] protected alias Matrix.zero_dotProduct' := zero_dotProduct' + @[simp] theorem add_dotProduct : (u + v) ⬝ᵥ w = u ⬝ᵥ w + v ⬝ᵥ w := by simp [dotProduct, add_mul, Finset.sum_add_distrib] +@[deprecated (since := "2024-12-12")] protected alias Matrix.add_dotProduct := add_dotProduct + @[simp] theorem dotProduct_add : u ⬝ᵥ (v + w) = u ⬝ᵥ v + u ⬝ᵥ w := by simp [dotProduct, mul_add, Finset.sum_add_distrib] +@[deprecated (since := "2024-12-12")] protected alias Matrix.dotProduct_add := dotProduct_add + @[simp] theorem sum_elim_dotProduct_sum_elim : Sum.elim u x ⬝ᵥ Sum.elim v y = u ⬝ᵥ v + x ⬝ᵥ y := by simp [dotProduct] +@[deprecated (since := "2024-12-12")] +protected alias Matrix.sum_elim_dotProduct_sum_elim := sum_elim_dotProduct_sum_elim + /-- Permuting a vector on the left of a dot product can be transferred to the right. -/ @[simp] theorem comp_equiv_symm_dotProduct (e : m ≃ n) : u ∘ e.symm ⬝ᵥ x = u ⬝ᵥ x ∘ e := (e.sum_comp _).symm.trans <| Finset.sum_congr rfl fun _ _ => by simp only [Function.comp, Equiv.symm_apply_apply] +@[deprecated (since := "2024-12-12")] +protected alias Matrix.comp_equiv_symm_dotProduct := comp_equiv_symm_dotProduct + /-- Permuting a vector on the right of a dot product can be transferred to the left. -/ @[simp] theorem dotProduct_comp_equiv_symm (e : n ≃ m) : u ⬝ᵥ x ∘ e.symm = u ∘ e ⬝ᵥ x := by simpa only [Equiv.symm_symm] using (comp_equiv_symm_dotProduct u x e.symm).symm +@[deprecated (since := "2024-12-12")] +protected alias Matrix.dotProduct_comp_equiv_symm := dotProduct_comp_equiv_symm + /-- Permuting vectors on both sides of a dot product is a no-op. -/ @[simp] theorem comp_equiv_dotProduct_comp_equiv (e : m ≃ n) : x ∘ e ⬝ᵥ y ∘ e = x ⬝ᵥ y := by -- Porting note: was `simp only` with all three lemmas rw [← dotProduct_comp_equiv_symm]; simp only [Function.comp_def, Equiv.apply_symm_apply] +@[deprecated (since := "2024-12-12")] +protected alias Matrix.comp_equiv_dotProduct_comp_equiv := comp_equiv_dotProduct_comp_equiv + end NonUnitalNonAssocSemiring section NonUnitalNonAssocSemiringDecidable @@ -151,18 +185,27 @@ theorem diagonal_dotProduct (i : m) : diagonal v i ⬝ᵥ w = v i * w i := by simp [diagonal_apply_ne' _ hij] convert Finset.sum_eq_single i (fun j _ => this j) _ using 1 <;> simp +@[deprecated (since := "2024-12-12")] +protected alias Matrix.diagonal_dotProduct := diagonal_dotProduct + @[simp] theorem dotProduct_diagonal (i : m) : v ⬝ᵥ diagonal w i = v i * w i := by have : ∀ j ≠ i, v j * diagonal w i j = 0 := fun j hij => by simp [diagonal_apply_ne' _ hij] convert Finset.sum_eq_single i (fun j _ => this j) _ using 1 <;> simp +@[deprecated (since := "2024-12-12")] +protected alias Matrix.dotProduct_diagonal := dotProduct_diagonal + @[simp] theorem dotProduct_diagonal' (i : m) : (v ⬝ᵥ fun j => diagonal w j i) = v i * w i := by have : ∀ j ≠ i, v j * diagonal w j i = 0 := fun j hij => by simp [diagonal_apply_ne _ hij] convert Finset.sum_eq_single i (fun j _ => this j) _ using 1 <;> simp +@[deprecated (since := "2024-12-12")] +protected alias Matrix.dotProduct_diagonal' := dotProduct_diagonal' + @[simp] theorem single_dotProduct (x : α) (i : m) : Pi.single i x ⬝ᵥ v = x * v i := by -- Porting note: (implicit arg) added `(f := fun _ => α)` @@ -170,6 +213,8 @@ theorem single_dotProduct (x : α) (i : m) : Pi.single i x ⬝ᵥ v = x * v i := simp [Pi.single_eq_of_ne hij] convert Finset.sum_eq_single i (fun j _ => this j) _ using 1 <;> simp +@[deprecated (since := "2024-12-12")] protected alias Matrix.single_dotProduct := single_dotProduct + @[simp] theorem dotProduct_single (x : α) (i : m) : v ⬝ᵥ Pi.single i x = v i * x := by -- Porting note: (implicit arg) added `(f := fun _ => α)` @@ -177,6 +222,8 @@ theorem dotProduct_single (x : α) (i : m) : v ⬝ᵥ Pi.single i x = v i * x := simp [Pi.single_eq_of_ne hij] convert Finset.sum_eq_single i (fun j _ => this j) _ using 1 <;> simp +@[deprecated (since := "2024-12-12")] protected alias Matrix.dotProduct_single := dotProduct_single + end NonUnitalNonAssocSemiringDecidable section NonAssocSemiring @@ -187,14 +234,23 @@ variable [NonAssocSemiring α] theorem one_dotProduct_one : (1 : n → α) ⬝ᵥ 1 = Fintype.card n := by simp [dotProduct] +@[deprecated (since := "2024-12-12")] +protected alias Matrix.one_dotProduct_one := one_dotProduct_one + theorem dotProduct_single_one [DecidableEq n] (v : n → α) (i : n) : dotProduct v (Pi.single i 1) = v i := by rw [dotProduct_single, mul_one] +@[deprecated (since := "2024-12-12")] +protected alias Matrix.dotProduct_single_one := dotProduct_single_one + theorem single_one_dotProduct [DecidableEq n] (i : n) (v : n → α) : dotProduct (Pi.single i 1) v = v i := by rw [single_dotProduct, one_mul] +@[deprecated (since := "2024-12-12")] +protected alias Matrix.single_one_dotProduct := single_one_dotProduct + end NonAssocSemiring section NonUnitalNonAssocRing @@ -204,18 +260,29 @@ variable [NonUnitalNonAssocRing α] (u v w : m → α) @[simp] theorem neg_dotProduct : -v ⬝ᵥ w = -(v ⬝ᵥ w) := by simp [dotProduct] +@[deprecated (since := "2024-12-12")] protected alias Matrix.neg_dotProduct := neg_dotProduct + @[simp] theorem dotProduct_neg : v ⬝ᵥ -w = -(v ⬝ᵥ w) := by simp [dotProduct] +@[deprecated (since := "2024-12-12")] protected alias Matrix.dotProduct_neg := dotProduct_neg + lemma neg_dotProduct_neg : -v ⬝ᵥ -w = v ⬝ᵥ w := by rw [neg_dotProduct, dotProduct_neg, neg_neg] +@[deprecated (since := "2024-12-12")] +protected alias Matrix.neg_dotProduct_neg := neg_dotProduct_neg + @[simp] theorem sub_dotProduct : (u - v) ⬝ᵥ w = u ⬝ᵥ w - v ⬝ᵥ w := by simp [sub_eq_add_neg] +@[deprecated (since := "2024-12-12")] protected alias Matrix.sub_dotProduct := sub_dotProduct + @[simp] theorem dotProduct_sub : u ⬝ᵥ (v - w) = u ⬝ᵥ v - u ⬝ᵥ w := by simp [sub_eq_add_neg] +@[deprecated (since := "2024-12-12")] protected alias Matrix.dotProduct_sub := dotProduct_sub + end NonUnitalNonAssocRing section DistribMulAction @@ -226,16 +293,22 @@ variable [Monoid R] [Mul α] [AddCommMonoid α] [DistribMulAction R α] theorem smul_dotProduct [IsScalarTower R α α] (x : R) (v w : m → α) : x • v ⬝ᵥ w = x • (v ⬝ᵥ w) := by simp [dotProduct, Finset.smul_sum, smul_mul_assoc] +@[deprecated (since := "2024-12-12")] protected alias Matrix.smul_dotProduct := smul_dotProduct + @[simp] theorem dotProduct_smul [SMulCommClass R α α] (x : R) (v w : m → α) : v ⬝ᵥ x • w = x • (v ⬝ᵥ w) := by simp [dotProduct, Finset.smul_sum, mul_smul_comm] +@[deprecated (since := "2024-12-12")] protected alias Matrix.dotProduct_smul := dotProduct_smul + end DistribMulAction end DotProduct open Matrix +namespace Matrix + /-- `M * N` is the usual product of matrices `M` and `N`, i.e. we have that `(M * N) i k` is the dot product of the `i`-th row of `M` by the `k`-th column of `N`. This is currently only defined when `m` is finite. -/ @@ -531,7 +604,7 @@ variable [NonUnitalNonAssocSemiring α] where `v` is seen as a column vector. Put another way, `M *ᵥ v` is the vector whose entries are those of `M * col v` (see `col_mulVec`). -The notation has precedence 73, which comes immediately before ` ⬝ᵥ ` for `Matrix.dotProduct`, +The notation has precedence 73, which comes immediately before ` ⬝ᵥ ` for `dotProduct`, so that `A *ᵥ v ⬝ᵥ B *ᵥ w` is parsed as `(A *ᵥ v) ⬝ᵥ (B *ᵥ w)`. -/ def mulVec [Fintype n] (M : Matrix m n α) (v : n → α) : m → α @@ -545,7 +618,7 @@ scoped infixr:73 " *ᵥ " => Matrix.mulVec where `v` is seen as a row vector. Put another way, `v ᵥ* M` is the vector whose entries are those of `row v * M` (see `row_vecMul`). -The notation has precedence 73, which comes immediately before ` ⬝ᵥ ` for `Matrix.dotProduct`, +The notation has precedence 73, which comes immediately before ` ⬝ᵥ ` for `dotProduct`, so that `v ᵥ* A ⬝ᵥ w ᵥ* B` is parsed as `(v ᵥ* A) ⬝ᵥ (w ᵥ* B)`. -/ def vecMul [Fintype m] (v : m → α) (M : Matrix m n α) : n → α @@ -935,7 +1008,7 @@ theorem map_matrix_mul (M : Matrix m n α) (N : Matrix n o α) (i : m) (j : o) ( theorem map_dotProduct [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) (v w : n → R) : f (v ⬝ᵥ w) = f ∘ v ⬝ᵥ f ∘ w := by - simp only [Matrix.dotProduct, map_sum f, f.map_mul, Function.comp] + simp only [dotProduct, map_sum f, f.map_mul, Function.comp] theorem map_vecMul [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) (M : Matrix n m R) (v : n → R) (i : m) : f ((v ᵥ* M) i) = ((f ∘ v) ᵥ* M.map f) i := by diff --git a/Mathlib/Data/Matrix/Reflection.lean b/Mathlib/Data/Matrix/Reflection.lean index a2339cbc23f8c..4789ae9f2be93 100644 --- a/Mathlib/Data/Matrix/Reflection.lean +++ b/Mathlib/Data/Matrix/Reflection.lean @@ -22,7 +22,7 @@ corresponding `*_eq` lemmas to be used in a place where they are definitionally ## Main definitions * `Matrix.transposeᵣ` -* `Matrix.dotProductᵣ` +* `dotProductᵣ` * `Matrix.mulᵣ` * `Matrix.mulVecᵣ` * `Matrix.vecMulᵣ` @@ -107,7 +107,7 @@ theorem transposeᵣ_eq : ∀ {m n} (A : Matrix (Fin m) (Fin n) α), transpose example (a b c d : α) : transpose !![a, b; c, d] = !![a, c; b, d] := (transposeᵣ_eq _).symm -/-- `Matrix.dotProduct` with better defeq for `Fin` -/ +/-- `dotProduct` with better defeq for `Fin` -/ def dotProductᵣ [Mul α] [Add α] [Zero α] {m} (a b : Fin m → α) : α := FinVec.sum <| FinVec.seq (FinVec.map (· * ·) a) b diff --git a/Mathlib/LinearAlgebra/AffineSpace/Matrix.lean b/Mathlib/LinearAlgebra/AffineSpace/Matrix.lean index 3ed1a0ba08f47..32b9394428889 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Matrix.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Matrix.lean @@ -108,7 +108,7 @@ theorem toMatrix_vecMul_coords (x : P) : b₂.coords x ᵥ* b.toMatrix b₂ = b. change _ = b.coord j x conv_rhs => rw [← b₂.affineCombination_coord_eq_self x] rw [Finset.map_affineCombination _ _ _ (b₂.sum_coord_apply_eq_one x)] - simp [Matrix.vecMul, Matrix.dotProduct, toMatrix_apply, coords] + simp [Matrix.vecMul, dotProduct, toMatrix_apply, coords] variable [DecidableEq ι] diff --git a/Mathlib/LinearAlgebra/CrossProduct.lean b/Mathlib/LinearAlgebra/CrossProduct.lean index 3093c7a58ad50..913de0b9539bc 100644 --- a/Mathlib/LinearAlgebra/CrossProduct.lean +++ b/Mathlib/LinearAlgebra/CrossProduct.lean @@ -86,7 +86,7 @@ theorem dot_self_cross (v w : Fin 3 → R) : v ⬝ᵥ v ×₃ w = 0 := by /-- The cross product of two vectors is perpendicular to the second vector. -/ @[simp 1100] -- Porting note: increase priority so that the LHS doesn't simplify theorem dot_cross_self (v w : Fin 3 → R) : w ⬝ᵥ v ×₃ w = 0 := by - rw [← cross_anticomm, Matrix.dotProduct_neg, dot_self_cross, neg_zero] + rw [← cross_anticomm, dotProduct_neg, dot_self_cross, neg_zero] /-- Cyclic permutations preserve the triple product. See also `triple_product_eq_det`. -/ theorem triple_product_permutation (u v w : Fin 3 → R) : u ⬝ᵥ v ×₃ w = v ⬝ᵥ w ×₃ u := by @@ -140,7 +140,7 @@ theorem jacobi_cross (u v w : Fin 3 → R) : u ×₃ (v ×₃ w) + v ×₃ (w × end LeibnizProperties --- this can also be proved via `Matrix.dotProduct_eq_zero_iff` and `triple_product_eq_det`, but +-- this can also be proved via `dotProduct_eq_zero_iff` and `triple_product_eq_det`, but -- that would require much heavier imports. lemma crossProduct_ne_zero_iff_linearIndependent {F : Type*} [Field F] {v w : Fin 3 → F} : crossProduct v w ≠ 0 ↔ LinearIndependent F ![v, w] := by diff --git a/Mathlib/LinearAlgebra/Matrix/BilinearForm.lean b/Mathlib/LinearAlgebra/Matrix/BilinearForm.lean index 7c27921ccd099..af8b96d719607 100644 --- a/Mathlib/LinearAlgebra/Matrix/BilinearForm.lean +++ b/Mathlib/LinearAlgebra/Matrix/BilinearForm.lean @@ -100,7 +100,7 @@ theorem Matrix.toBilin'_apply (M : Matrix n n R₁) (x y : n → R₁) : (by simp only [smul_eq_mul, mul_assoc, mul_comm, mul_left_comm]) theorem Matrix.toBilin'_apply' (M : Matrix n n R₁) (v w : n → R₁) : - Matrix.toBilin' M v w = Matrix.dotProduct v (M *ᵥ w) := Matrix.toLinearMap₂'_apply' _ _ _ + Matrix.toBilin' M v w = dotProduct v (M *ᵥ w) := Matrix.toLinearMap₂'_apply' _ _ _ @[simp] theorem Matrix.toBilin'_single (M : Matrix n n R₁) (i j : n) : diff --git a/Mathlib/LinearAlgebra/Matrix/DotProduct.lean b/Mathlib/LinearAlgebra/Matrix/DotProduct.lean index 1a249d247e62f..0d151dcdfbceb 100644 --- a/Mathlib/LinearAlgebra/Matrix/DotProduct.lean +++ b/Mathlib/LinearAlgebra/Matrix/DotProduct.lean @@ -10,14 +10,14 @@ import Mathlib.LinearAlgebra.StdBasis /-! # Dot product of two vectors -This file contains some results on the map `Matrix.dotProduct`, which maps two +This file contains some results on the map `dotProduct`, which maps two vectors `v w : n → R` to the sum of the entrywise products `v i * w i`. ## Main results -* `Matrix.dotProduct_stdBasis_one`: the dot product of `v` with the `i`th +* `dotProduct_stdBasis_one`: the dot product of `v` with the `i`th standard basis vector is `v i` -* `Matrix.dotProduct_eq_zero_iff`: if `v`'s dot product with all `w` is zero, +* `dotProduct_eq_zero_iff`: if `v`'s dot product with all `w` is zero, then `v` is zero ## Tags @@ -29,20 +29,18 @@ matrix, reindex variable {m n p R : Type*} -namespace Matrix - section Semiring variable [Semiring R] [Fintype n] set_option linter.deprecated false in -@[simp, deprecated dotProduct_single (since := "2024-08-09")] +@[simp, deprecated Matrix.dotProduct_single (since := "2024-08-09")] theorem dotProduct_stdBasis_eq_mul [DecidableEq n] (v : n → R) (c : R) (i : n) : dotProduct v (LinearMap.stdBasis R (fun _ => R) i c) = v i * c := dotProduct_single .. set_option linter.deprecated false in -@[deprecated dotProduct_single_one (since := "2024-08-09")] +@[deprecated Matrix.dotProduct_single_one (since := "2024-08-09")] theorem dotProduct_stdBasis_one [DecidableEq n] (v : n → R) (i : n) : dotProduct v (LinearMap.stdBasis R (fun _ => R) i 1) = v i := dotProduct_single_one .. @@ -51,15 +49,25 @@ theorem dotProduct_eq (v w : n → R) (h : ∀ u, dotProduct v u = dotProduct w funext x classical rw [← dotProduct_single_one v x, ← dotProduct_single_one w x, h] +@[deprecated (since := "2024-12-12")] protected alias Matrix.dotProduct_eq := dotProduct_eq + theorem dotProduct_eq_iff {v w : n → R} : (∀ u, dotProduct v u = dotProduct w u) ↔ v = w := ⟨fun h => dotProduct_eq v w h, fun h _ => h ▸ rfl⟩ +@[deprecated (since := "2024-12-12")] protected alias Matrix.dotProduct_eq_iff := dotProduct_eq_iff + theorem dotProduct_eq_zero (v : n → R) (h : ∀ w, dotProduct v w = 0) : v = 0 := dotProduct_eq _ _ fun u => (h u).symm ▸ (zero_dotProduct u).symm +@[deprecated (since := "2024-12-12")] +protected alias Matrix.dotProduct_eq_zero := dotProduct_eq_zero + theorem dotProduct_eq_zero_iff {v : n → R} : (∀ w, dotProduct v w = 0) ↔ v = 0 := ⟨fun h => dotProduct_eq_zero v h, fun h w => h.symm ▸ zero_dotProduct w⟩ +@[deprecated (since := "2024-12-12")] +protected alias Matrix.dotProduct_eq_zero_iff := dotProduct_eq_zero_iff + end Semiring section OrderedSemiring @@ -69,14 +77,25 @@ variable [OrderedSemiring R] [Fintype n] lemma dotProduct_nonneg_of_nonneg {v w : n → R} (hv : 0 ≤ v) (hw : 0 ≤ w) : 0 ≤ dotProduct v w := Finset.sum_nonneg (fun i _ => mul_nonneg (hv i) (hw i)) +@[deprecated (since := "2024-12-12")] +protected alias Matrix.dotProduct_nonneg_of_nonneg := dotProduct_nonneg_of_nonneg + lemma dotProduct_le_dotProduct_of_nonneg_right {u v w : n → R} (huv : u ≤ v) (hw : 0 ≤ w) : dotProduct u w ≤ dotProduct v w := Finset.sum_le_sum (fun i _ => mul_le_mul_of_nonneg_right (huv i) (hw i)) +@[deprecated (since := "2024-12-12")] +protected alias Matrix.dotProduct_le_dotProduct_of_nonneg_right := + dotProduct_le_dotProduct_of_nonneg_right + lemma dotProduct_le_dotProduct_of_nonneg_left {u v w : n → R} (huv : u ≤ v) (hw : 0 ≤ w) : dotProduct w u ≤ dotProduct w v := Finset.sum_le_sum (fun i _ => mul_le_mul_of_nonneg_left (huv i) (hw i)) +@[deprecated (since := "2024-12-12")] +protected alias Matrix.dotProduct_le_dotProduct_of_nonneg_left := + dotProduct_le_dotProduct_of_nonneg_left + end OrderedSemiring section Self @@ -88,6 +107,9 @@ theorem dotProduct_self_eq_zero [LinearOrderedRing R] {v : n → R} : dotProduct (Finset.sum_eq_zero_iff_of_nonneg fun i _ => mul_self_nonneg (v i)).trans <| by simp [funext_iff] +@[deprecated (since := "2024-12-12")] +protected alias Matrix.dotProduct_self_eq_zero := dotProduct_self_eq_zero + section StarOrderedRing variable [PartialOrder R] [NonUnitalRing R] [StarRing R] [StarOrderedRing R] @@ -97,11 +119,17 @@ variable [PartialOrder R] [NonUnitalRing R] [StarRing R] [StarOrderedRing R] theorem dotProduct_star_self_nonneg (v : n → R) : 0 ≤ dotProduct (star v) v := Fintype.sum_nonneg fun _ => star_mul_self_nonneg _ +@[deprecated (since := "2024-12-12")] +protected alias Matrix.dotProduct_star_self_nonneg := dotProduct_star_self_nonneg + /-- Note that this applies to `ℂ` via `RCLike.toStarOrderedRing`. -/ @[simp] theorem dotProduct_self_star_nonneg (v : n → R) : 0 ≤ dotProduct v (star v) := Fintype.sum_nonneg fun _ => mul_star_self_nonneg _ +@[deprecated (since := "2024-12-12")] +protected alias Matrix.dotProduct_self_star_nonneg := dotProduct_self_star_nonneg + variable [NoZeroDivisors R] /-- Note that this applies to `ℂ` via `RCLike.toStarOrderedRing`. -/ @@ -110,12 +138,20 @@ theorem dotProduct_star_self_eq_zero {v : n → R} : dotProduct (star v) v = 0 (Fintype.sum_eq_zero_iff_of_nonneg fun _ => star_mul_self_nonneg _).trans <| by simp [funext_iff, mul_eq_zero] +@[deprecated (since := "2024-12-12")] +protected alias Matrix.dotProduct_star_self_eq_zero := dotProduct_star_self_eq_zero + /-- Note that this applies to `ℂ` via `RCLike.toStarOrderedRing`. -/ @[simp] theorem dotProduct_self_star_eq_zero {v : n → R} : dotProduct v (star v) = 0 ↔ v = 0 := (Fintype.sum_eq_zero_iff_of_nonneg fun _ => mul_star_self_nonneg _).trans <| by simp [funext_iff, mul_eq_zero] +@[deprecated (since := "2024-12-12")] +protected alias Matrix.dotProduct_self_star_eq_zero := dotProduct_self_star_eq_zero + +namespace Matrix + @[simp] lemma conjTranspose_mul_self_eq_zero {n} {A : Matrix m n R} : Aᴴ * A = 0 ↔ A = 0 := ⟨fun h => Matrix.ext fun i j => @@ -186,8 +222,8 @@ theorem dotProduct_star_self_pos_iff {v : n → R} : theorem dotProduct_self_star_pos_iff {v : n → R} : 0 < dotProduct v (star v) ↔ v ≠ 0 := by simpa using dotProduct_star_self_pos_iff (v := star v) +end Matrix + end StarOrderedRing end Self - -end Matrix diff --git a/Mathlib/LinearAlgebra/Matrix/Gershgorin.lean b/Mathlib/LinearAlgebra/Matrix/Gershgorin.lean index 06f6a65c35c43..ca6059ced1fbb 100644 --- a/Mathlib/LinearAlgebra/Matrix/Gershgorin.lean +++ b/Mathlib/LinearAlgebra/Matrix/Gershgorin.lean @@ -45,7 +45,7 @@ theorem eigenvalue_mem_ball {μ : K} (hμ : Module.End.HasEigenvalue (Matrix.toL _ = ‖(A i i * v i - μ * v i) * (v i)⁻¹‖ := by congr; field_simp [h_nz]; ring _ = ‖(A i i * v i - ∑ j, A i j * v j) * (v i)⁻¹‖ := by rw [show μ * v i = ∑ x : n, A i x * v x by - rw [← Matrix.dotProduct, ← Matrix.mulVec] + rw [← dotProduct, ← Matrix.mulVec] exact (congrFun (Module.End.mem_eigenspace_iff.mp h_eg) i).symm] _ = ‖(∑ j ∈ Finset.univ.erase i, A i j * v j) * (v i)⁻¹‖ := by rw [Finset.sum_erase_eq_sub (Finset.mem_univ i), ← neg_sub, neg_mul, norm_neg] diff --git a/Mathlib/LinearAlgebra/Matrix/Nondegenerate.lean b/Mathlib/LinearAlgebra/Matrix/Nondegenerate.lean index 9b9d6afa52ed7..579889da53fc4 100644 --- a/Mathlib/LinearAlgebra/Matrix/Nondegenerate.lean +++ b/Mathlib/LinearAlgebra/Matrix/Nondegenerate.lean @@ -24,16 +24,16 @@ variable {m R A : Type*} [Fintype m] [CommRing R] /-- A matrix `M` is nondegenerate if for all `v ≠ 0`, there is a `w ≠ 0` with `w * M * v ≠ 0`. -/ def Nondegenerate (M : Matrix m m R) := - ∀ v, (∀ w, Matrix.dotProduct v (M *ᵥ w) = 0) → v = 0 + ∀ v, (∀ w, dotProduct v (M *ᵥ w) = 0) → v = 0 /-- If `M` is nondegenerate and `w * M * v = 0` for all `w`, then `v = 0`. -/ theorem Nondegenerate.eq_zero_of_ortho {M : Matrix m m R} (hM : Nondegenerate M) {v : m → R} - (hv : ∀ w, Matrix.dotProduct v (M *ᵥ w) = 0) : v = 0 := + (hv : ∀ w, dotProduct v (M *ᵥ w) = 0) : v = 0 := hM v hv /-- If `M` is nondegenerate and `v ≠ 0`, then there is some `w` such that `w * M * v ≠ 0`. -/ theorem Nondegenerate.exists_not_ortho_of_ne_zero {M : Matrix m m R} (hM : Nondegenerate M) - {v : m → R} (hv : v ≠ 0) : ∃ w, Matrix.dotProduct v (M *ᵥ w) ≠ 0 := + {v : m → R} (hv : v ≠ 0) : ∃ w, dotProduct v (M *ᵥ w) ≠ 0 := not_forall.mp (mt hM.eq_zero_of_ortho hv) variable [CommRing A] [IsDomain A] diff --git a/Mathlib/LinearAlgebra/Matrix/Orthogonal.lean b/Mathlib/LinearAlgebra/Matrix/Orthogonal.lean index b5cfe6c76c4bb..774a61bed869a 100644 --- a/Mathlib/LinearAlgebra/Matrix/Orthogonal.lean +++ b/Mathlib/LinearAlgebra/Matrix/Orthogonal.lean @@ -32,12 +32,12 @@ variable (A : Matrix m n α) open Matrix /-- `A.HasOrthogonalRows` means matrix `A` has orthogonal rows (with respect to -`Matrix.dotProduct`). -/ +`dotProduct`). -/ def HasOrthogonalRows [Fintype n] : Prop := ∀ ⦃i₁ i₂⦄, i₁ ≠ i₂ → dotProduct (A i₁) (A i₂) = 0 /-- `A.HasOrthogonalCols` means matrix `A` has orthogonal columns (with respect to -`Matrix.dotProduct`). -/ +`dotProduct`). -/ def HasOrthogonalCols [Fintype m] : Prop := HasOrthogonalRows Aᵀ diff --git a/Mathlib/LinearAlgebra/Matrix/PosDef.lean b/Mathlib/LinearAlgebra/Matrix/PosDef.lean index 0f7ca95450e7f..fb74430f25779 100644 --- a/Mathlib/LinearAlgebra/Matrix/PosDef.lean +++ b/Mathlib/LinearAlgebra/Matrix/PosDef.lean @@ -90,7 +90,7 @@ theorem submatrix {M : Matrix n n R} (hM : M.PosSemidef) (e : m → n) : theorem transpose {M : Matrix n n R} (hM : M.PosSemidef) : Mᵀ.PosSemidef := by refine ⟨IsHermitian.transpose hM.1, fun x => ?_⟩ convert hM.2 (star x) using 1 - rw [mulVec_transpose, Matrix.dotProduct_mulVec, star_star, dotProduct_comm] + rw [mulVec_transpose, dotProduct_mulVec, star_star, dotProduct_comm] theorem conjTranspose {M : Matrix n n R} (hM : M.PosSemidef) : Mᴴ.PosSemidef := hM.1.symm ▸ hM @@ -243,7 +243,7 @@ lemma eq_of_sq_eq_sq {B : Matrix n n 𝕜} (hB : PosSemidef B) (hAB : A ^ 2 = B have aux : star v ⬝ᵥ (A - B) *ᵥ v = 0 := by rw [sub_mulVec, dotProduct_sub, h_van.1, h_van.2, sub_zero] rw [hv', dotProduct_smul, RCLike.real_smul_eq_coe_mul, ← mul_zero ↑t] at aux - exact hv <| Matrix.dotProduct_star_self_eq_zero.mp <| mul_left_cancel₀ + exact hv <| dotProduct_star_self_eq_zero.mp <| mul_left_cancel₀ (RCLike.ofReal_ne_zero.mpr ht) aux lemma sqrt_sq : (hA.pow 2 : PosSemidef (A ^ 2)).sqrt = A := @@ -345,7 +345,7 @@ theorem posSemidef {M : Matrix n n R} (hM : M.PosDef) : M.PosSemidef := by theorem transpose {M : Matrix n n R} (hM : M.PosDef) : Mᵀ.PosDef := by refine ⟨IsHermitian.transpose hM.1, fun x hx => ?_⟩ convert hM.2 (star x) (star_ne_zero.2 hx) using 1 - rw [mulVec_transpose, Matrix.dotProduct_mulVec, star_star, dotProduct_comm] + rw [mulVec_transpose, dotProduct_mulVec, star_star, dotProduct_comm] protected theorem diagonal [StarOrderedRing R] [DecidableEq n] [NoZeroDivisors R] {d : n → R} (h : ∀ i, 0 < d i) : diff --git a/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean b/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean index 150f3e2f40e25..d6a0f32b76bd3 100644 --- a/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean +++ b/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean @@ -185,8 +185,8 @@ theorem Matrix.toLinearMap₂'_apply (M : Matrix n m N₂) (x : n → S₁) (y : rw [RingHom.id_apply, RingHom.id_apply, smul_comm] theorem Matrix.toLinearMap₂'_apply' {T : Type*} [CommSemiring T] (M : Matrix n m T) (v : n → T) - (w : m → T) : Matrix.toLinearMap₂' T M v w = Matrix.dotProduct v (M *ᵥ w) := by - simp_rw [Matrix.toLinearMap₂'_apply, Matrix.dotProduct, Matrix.mulVec, Matrix.dotProduct] + (w : m → T) : Matrix.toLinearMap₂' T M v w = dotProduct v (M *ᵥ w) := by + simp_rw [Matrix.toLinearMap₂'_apply, dotProduct, Matrix.mulVec, dotProduct] refine Finset.sum_congr rfl fun _ _ => ?_ rw [Finset.mul_sum] refine Finset.sum_congr rfl fun _ _ => ?_ diff --git a/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean b/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean index a3dfd72c6e079..558dc8dc12455 100644 --- a/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean +++ b/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean @@ -475,7 +475,7 @@ theorem coe_T_zpow (n : ℤ) : (T ^ n).1 = !![1, n; 0, 1] := by @[simp] theorem T_pow_mul_apply_one (n : ℤ) (g : SL(2, ℤ)) : (T ^ n * g) 1 = g 1 := by ext j - simp [coe_T_zpow, Matrix.vecMul, Matrix.dotProduct, Fin.sum_univ_succ, vecTail] + simp [coe_T_zpow, Matrix.vecMul, dotProduct, Fin.sum_univ_succ, vecTail] @[simp] theorem T_mul_apply_one (g : SL(2, ℤ)) : (T * g) 1 = g 1 := by diff --git a/Mathlib/LinearAlgebra/Matrix/Spectrum.lean b/Mathlib/LinearAlgebra/Matrix/Spectrum.lean index 9f15fd930b639..3367f16ef2d73 100644 --- a/Mathlib/LinearAlgebra/Matrix/Spectrum.lean +++ b/Mathlib/LinearAlgebra/Matrix/Spectrum.lean @@ -115,7 +115,7 @@ theorem spectral_theorem : ← mul_assoc, (Matrix.mem_unitaryGroup_iff).mp (eigenvectorUnitary hA).2, one_mul] theorem eigenvalues_eq (i : n) : - (hA.eigenvalues i) = RCLike.re (Matrix.dotProduct (star ⇑(hA.eigenvectorBasis i)) + (hA.eigenvalues i) = RCLike.re (dotProduct (star ⇑(hA.eigenvectorBasis i)) (A *ᵥ ⇑(hA.eigenvectorBasis i))) := by simp only [mulVec_eigenvectorBasis, dotProduct_smul,← EuclideanSpace.inner_eq_star_dotProduct, inner_self_eq_norm_sq_to_K, RCLike.smul_re, hA.eigenvectorBasis.orthonormal.1 i, diff --git a/Mathlib/LinearAlgebra/Matrix/ToLin.lean b/Mathlib/LinearAlgebra/Matrix/ToLin.lean index 1ef1d74c925eb..6bc25e1d2b4ec 100644 --- a/Mathlib/LinearAlgebra/Matrix/ToLin.lean +++ b/Mathlib/LinearAlgebra/Matrix/ToLin.lean @@ -765,7 +765,7 @@ theorem Matrix.toLinAlgEquiv_mul (A B : Matrix n n R) : theorem Matrix.toLin_finTwoProd_apply (a b c d : R) (x : R × R) : Matrix.toLin (Basis.finTwoProd R) (Basis.finTwoProd R) !![a, b; c, d] x = (a * x.fst + b * x.snd, c * x.fst + d * x.snd) := by - simp [Matrix.toLin_apply, Matrix.mulVec, Matrix.dotProduct] + simp [Matrix.toLin_apply, Matrix.mulVec, dotProduct] theorem Matrix.toLin_finTwoProd (a b c d : R) : Matrix.toLin (Basis.finTwoProd R) (Basis.finTwoProd R) !![a, b; c, d] = diff --git a/Mathlib/LinearAlgebra/Projectivization/Constructions.lean b/Mathlib/LinearAlgebra/Projectivization/Constructions.lean index cdbb83b605a70..3572826fa6f6e 100644 --- a/Mathlib/LinearAlgebra/Projectivization/Constructions.lean +++ b/Mathlib/LinearAlgebra/Projectivization/Constructions.lean @@ -31,24 +31,24 @@ section DotProduct /-- Orthogonality on the projective plane. -/ def orthogonal : ℙ F (m → F) → ℙ F (m → F) → Prop := - Quotient.lift₂ (fun v w ↦ Matrix.dotProduct v.1 w.1 = 0) (fun _ _ _ _ ⟨_, h1⟩ ⟨_, h2⟩ ↦ by - simp_rw [← h1, ← h2, Matrix.dotProduct_smul, Matrix.smul_dotProduct, smul_smul, + Quotient.lift₂ (fun v w ↦ dotProduct v.1 w.1 = 0) (fun _ _ _ _ ⟨_, h1⟩ ⟨_, h2⟩ ↦ by + simp_rw [← h1, ← h2, dotProduct_smul, smul_dotProduct, smul_smul, smul_eq_zero_iff_eq]) lemma orthogonal_mk {v w : m → F} (hv : v ≠ 0) (hw : w ≠ 0) : - orthogonal (mk F v hv) (mk F w hw) ↔ Matrix.dotProduct v w = 0 := + orthogonal (mk F v hv) (mk F w hw) ↔ dotProduct v w = 0 := Iff.rfl lemma orthogonal_comm {v w : ℙ F (m → F)} : orthogonal v w ↔ orthogonal w v := by induction' v with v hv induction' w with w hw - rw [orthogonal_mk hv hw, orthogonal_mk hw hv, Matrix.dotProduct_comm] + rw [orthogonal_mk hv hw, orthogonal_mk hw hv, dotProduct_comm] lemma exists_not_self_orthogonal (v : ℙ F (m → F)) : ∃ w, ¬ orthogonal v w := by induction' v with v hv - rw [ne_eq, ← Matrix.dotProduct_eq_zero_iff, not_forall] at hv + rw [ne_eq, ← dotProduct_eq_zero_iff, not_forall] at hv obtain ⟨w, hw⟩ := hv - exact ⟨mk F w fun h ↦ hw (by rw [h, Matrix.dotProduct_zero]), hw⟩ + exact ⟨mk F w fun h ↦ hw (by rw [h, dotProduct_zero]), hw⟩ lemma exists_not_orthogonal_self (v : ℙ F (m → F)) : ∃ w, ¬ orthogonal w v := by simp only [orthogonal_comm] @@ -113,7 +113,7 @@ theorem cross_orthogonal_left {v w : ℙ F (Fin 3 → F)} (h : v ≠ w) : (cross v w).orthogonal v := by induction' v with v hv induction' w with w hw - rw [cross_mk_of_ne hv hw h, orthogonal_mk, Matrix.dotProduct_comm, dot_self_cross] + rw [cross_mk_of_ne hv hw h, orthogonal_mk, dotProduct_comm, dot_self_cross] theorem cross_orthogonal_right {v w : ℙ F (Fin 3 → F)} (h : v ≠ w) : (cross v w).orthogonal w := by diff --git a/Mathlib/Tactic/ScopedNS.lean b/Mathlib/Tactic/ScopedNS.lean index 0057e00d2ffe4..dd751fc5170b1 100644 --- a/Mathlib/Tactic/ScopedNS.lean +++ b/Mathlib/Tactic/ScopedNS.lean @@ -18,8 +18,8 @@ open Lean `scoped[NS]` is similar to the `scoped` modifier on attributes and notations, but it scopes the syntax in the specified namespace instead of the current namespace. ``` -scoped[Matrix] infixl:72 " ⬝ᵥ " => Matrix.dotProduct --- declares `*` as a notation for vector dot productss +scoped[Matrix] postfix:1024 "ᵀ" => Matrix.transpose +-- declares `ᵀ` as a notation for matrix transposition -- which is only accessible if you `open Matrix` or `open scoped Matrix`. namespace Nat diff --git a/scripts/nolints_prime_decls.txt b/scripts/nolints_prime_decls.txt index 3304424a06b19..47c01ac4b1533 100644 --- a/scripts/nolints_prime_decls.txt +++ b/scripts/nolints_prime_decls.txt @@ -2571,8 +2571,8 @@ Matrix.diagonal_mul_diagonal' Matrix.diagonal_natCast' Matrix.diagonal_ofNat' Matrix.diagonal_toLin' -Matrix.dotProduct_diagonal' -Matrix.dotProduct_zero' +dotProduct_diagonal' +dotProduct_zero' Matrix.empty_val' Matrix.exists_mulVec_eq_zero_iff' Matrix.exp_blockDiagonal' @@ -2615,7 +2615,7 @@ Matrix.trace_mul_cycle' Matrix.twoBlockTriangular_det' Matrix.vec2_dotProduct' Matrix.vec3_dotProduct' -Matrix.zero_dotProduct' +zero_dotProduct' Matrix.zpow_mul' Matroid.Base.exchange_base_of_indep' Matroid.base_restrict_iff' From d131f928eb56172a4193db76d31ed6165b30dcb9 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Thu, 19 Dec 2024 10:47:36 +0000 Subject: [PATCH 780/829] feat: generators for SL(2, Z) (#15851) We show that SL(2, Z) is generated by S and T. Co-authored-by: Chris Birkbeck --- .../Matrix/FixedDetMatrices.lean | 41 +++++++++++++++++++ 1 file changed, 41 insertions(+) diff --git a/Mathlib/LinearAlgebra/Matrix/FixedDetMatrices.lean b/Mathlib/LinearAlgebra/Matrix/FixedDetMatrices.lean index 0154b1ab6322f..7c9dbfd95a909 100644 --- a/Mathlib/LinearAlgebra/Matrix/FixedDetMatrices.lean +++ b/Mathlib/LinearAlgebra/Matrix/FixedDetMatrices.lean @@ -13,6 +13,8 @@ import Mathlib.Data.Int.Interval This file defines the type of matrices with fixed determinant `m` and proves some basic results about them. +We also prove that the subgroup of `SL(2,ℤ)` generated by `S` and `T` is the whole group. + Note: Some of this was done originally in Lean 3 in the kbb (https://github.com/kim-em/kbb/tree/master) repository, so credit to those authors. -/ @@ -242,6 +244,45 @@ theorem induction_on {C : Δ m → Prop} {A : Δ m} (hm : m ≠ 0) rw [← reduce_reduceStep hc] at hA simpa only [reduceStep, prop_red_S hS, prop_red_T_pow hS hT] using ih hA +/--The subgroup of `SL(2,ℤ)` generated by `S` and `T`. -/ +def S_T_subgroup := Subgroup.closure {S, T} + +lemma S_mem_S_T_subgroup : S ∈ S_T_subgroup := by + apply Subgroup.subset_closure + simp only [Set.mem_insert_iff, Set.mem_singleton_iff, true_or] + +lemma T_mem_S_T_subgroup : T ∈ S_T_subgroup := by + apply Subgroup.subset_closure + simp only [Set.mem_insert_iff, Set.mem_singleton_iff, or_true] + +lemma reps_one_id (A : FixedDetMatrix (Fin 2) ℤ 1) (a1 : A.1 1 0 = 0) (a4 : 0 < A.1 0 0) + (a6 : |A.1 0 1| < |(A.1 1 1)|) : A = (1 : SL(2, ℤ)) := by + have := Int.mul_eq_one_iff_eq_one_or_neg_one.mp (A_c_eq_zero a1) + ext i j + fin_cases i <;> fin_cases j <;> aesop + +lemma det_one_mem_S_T_subgroup (A : Δ 1) : A ∈ S_T_subgroup := by + induction A using (induction_on Int.one_ne_zero) with + | h0 A a1 a4 _ a6 => + rw [reps_one_id A a1 a4 a6] + exact Subgroup.one_mem S_T_subgroup + | hS B hb => + rw [smul_eq_mul] + apply Subgroup.mul_mem _ (S_mem_S_T_subgroup) (hb) + | hT B hb => + rw [smul_eq_mul] + apply Subgroup.mul_mem _ (T_mem_S_T_subgroup) (hb) + end IntegralFixedDetMatrices end FixedDetMatrices + +open MatrixGroups FixedDetMatrices + +section SL2Z_generators + +/-- `SL(2, ℤ)` is generated by `S` and `T`. -/ +lemma SpecialLinearGroup.SL2Z_generators : S_T_subgroup = ⊤ := by + refine (Subgroup.eq_top_iff' S_T_subgroup).mpr (fun _ => det_one_mem_S_T_subgroup _) + +end SL2Z_generators From 8bcb2f38ecc27e75dac4010560c0d9d8ace2d2e2 Mon Sep 17 00:00:00 2001 From: damiano Date: Thu, 19 Dec 2024 10:47:37 +0000 Subject: [PATCH 781/829] fix: add `geometric-group-theory` autolabel (#20061) A recent PR added a new dir that was not covered by the auto-label automation. This PR adds awareness of this dir in autolabel: the new label is `t-geometric-group-theory`. [Zulip discussion](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/Auto-label.20PR.20topics/near/489857735) --- scripts/autolabel.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/scripts/autolabel.lean b/scripts/autolabel.lean index c848bd6564ef5..5225051d5e56a 100644 --- a/scripts/autolabel.lean +++ b/scripts/autolabel.lean @@ -105,6 +105,8 @@ def mathlibLabels : Array Label := #[ { label := "t-dynamics" }, { label := "t-euclidean-geometry", dirs := #["Mathlib" / "Geometry" / "Euclidean"] }, + { label := "t-geometric-group-theory", + dirs := #["Mathlib" / "Geometry" / "Group"] }, { label := "t-linter", dirs := #["Mathlib" / "Tactic" / "Linter"] }, { label := "t-logic", From 32ba402fc2048299c9c0e153b72b4d4ea8aade19 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Thu, 19 Dec 2024 10:56:34 +0000 Subject: [PATCH 782/829] chore: update Mathlib dependencies 2024-12-19 (#20058) 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 09adba8286a84..2ad65f505f796 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -45,7 +45,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "43bcb1964528411e47bfa4edd0c87d1face1fce4", + "rev": "a4a08d92be3de00def5298059bf707c72dfd3c66", "name": "aesop", "manifestFile": "lake-manifest.json", "inputRev": "master", From f00d276c55a4e9afac91f01d6a6b7c24d9040017 Mon Sep 17 00:00:00 2001 From: Pim Otte Date: Thu, 19 Dec 2024 13:05:27 +0000 Subject: [PATCH 783/829] feat(Combinatorics/SimpleGraph): setup of universal vertices (#20023) Added the defintion of `universalVerts` and `deleteUniversalVerts` along with some initial results. This is in preparation for a proof of Tutte's theorem. Co-authored-by: Pim Otte --- Mathlib.lean | 1 + .../SimpleGraph/UniversalVerts.lean | 52 +++++++++++++++++++ Mathlib/Data/Set/Card.lean | 3 ++ Mathlib/SetTheory/Cardinal/Finite.lean | 4 ++ 4 files changed, 60 insertions(+) create mode 100644 Mathlib/Combinatorics/SimpleGraph/UniversalVerts.lean diff --git a/Mathlib.lean b/Mathlib.lean index 15e3f119d136e..ac8fe9eacb347 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2291,6 +2291,7 @@ import Mathlib.Combinatorics.SimpleGraph.Triangle.Counting import Mathlib.Combinatorics.SimpleGraph.Triangle.Removal import Mathlib.Combinatorics.SimpleGraph.Triangle.Tripartite import Mathlib.Combinatorics.SimpleGraph.Turan +import Mathlib.Combinatorics.SimpleGraph.UniversalVerts import Mathlib.Combinatorics.SimpleGraph.Walk import Mathlib.Combinatorics.Young.SemistandardTableau import Mathlib.Combinatorics.Young.YoungDiagram diff --git a/Mathlib/Combinatorics/SimpleGraph/UniversalVerts.lean b/Mathlib/Combinatorics/SimpleGraph/UniversalVerts.lean new file mode 100644 index 0000000000000..7c7f3b1cb0d0c --- /dev/null +++ b/Mathlib/Combinatorics/SimpleGraph/UniversalVerts.lean @@ -0,0 +1,52 @@ +/- +Copyright (c) 2024 Pim Otte. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Pim Otte +-/ +import Mathlib.Combinatorics.SimpleGraph.Clique +import Mathlib.Combinatorics.SimpleGraph.Matching + +/-! +# Universal Vertices + +This file defines the set of universal vertices: those vertices that are connected +to all others. In addition, it describes results when considering connected components +of the graph where universal vertices are deleted. This particular graph plays a role +in the proof of Tutte's Theorem. + +## Main definitions + +* `G.universalVerts` is the set of vertices that are connected to all other vertices. +* `G.deleteUniversalVerts` is the subgraph of `G` with the universal vertices removed. +-/ + +namespace SimpleGraph +variable {V : Type*} {G : SimpleGraph V} + +/-- +The set of vertices that are connected to all other vertices. +-/ +def universalVerts (G : SimpleGraph V) : Set V := {v : V | ∀ ⦃w⦄, v ≠ w → G.Adj w v} + +lemma isClique_universalVerts (G : SimpleGraph V) : G.IsClique G.universalVerts := + fun _ _ _ hy hxy ↦ hy hxy.symm + +/-- +The subgraph of `G` with the universal vertices removed. +-/ +@[simps!] +def deleteUniversalVerts (G : SimpleGraph V) : Subgraph G := + (⊤ : Subgraph G).deleteVerts G.universalVerts + +lemma Subgraph.IsMatching.exists_of_universalVerts [Fintype V] {s : Set V} + (h : Disjoint G.universalVerts s) (hc : s.ncard ≤ G.universalVerts.ncard) : + ∃ t ⊆ G.universalVerts, ∃ (M : Subgraph G), M.verts = s ∪ t ∧ M.IsMatching := by + obtain ⟨t, ht⟩ := Set.exists_subset_card_eq hc + refine ⟨t, ht.1, ?_⟩ + obtain ⟨f⟩ : Nonempty (s ≃ t) := by + rw [← Cardinal.eq, ← t.cast_ncard t.toFinite, ← s.cast_ncard s.toFinite, ht.2] + letI hd := Set.disjoint_of_subset_left ht.1 h + have hadj (v : s) : G.Adj v (f v) := ht.1 (f v).2 (hd.ne_of_mem (f v).2 v.2) + exact Subgraph.IsMatching.exists_of_disjoint_sets_of_equiv hd.symm f hadj + +end SimpleGraph diff --git a/Mathlib/Data/Set/Card.lean b/Mathlib/Data/Set/Card.lean index 56ac372f3623f..d5b5ef3a80dc0 100644 --- a/Mathlib/Data/Set/Card.lean +++ b/Mathlib/Data/Set/Card.lean @@ -492,6 +492,9 @@ theorem ncard_eq_toFinset_card' (s : Set α) [Fintype s] : s.ncard = s.toFinset.card := by simp [← Nat.card_coe_set_eq, Nat.card_eq_fintype_card] +lemma cast_ncard {s : Set α} (hs : s.Finite) : + (s.ncard : Cardinal) = Cardinal.mk s := @Nat.cast_card _ hs + theorem encard_le_coe_iff_finite_ncard_le {k : ℕ} : s.encard ≤ k ↔ s.Finite ∧ s.ncard ≤ k := by rw [encard_le_coe_iff, and_congr_right_iff] exact fun hfin ↦ ⟨fun ⟨n₀, hn₀, hle⟩ ↦ by rwa [ncard_def, hn₀, ENat.toNat_coe], diff --git a/Mathlib/SetTheory/Cardinal/Finite.lean b/Mathlib/SetTheory/Cardinal/Finite.lean index 5de36466bcce0..afde38acbd05c 100644 --- a/Mathlib/SetTheory/Cardinal/Finite.lean +++ b/Mathlib/SetTheory/Cardinal/Finite.lean @@ -58,6 +58,10 @@ lemma card_eq_card_finite_toFinset {s : Set α} (hs : s.Finite) : Nat.card s = h @[simp] lemma card_eq_zero_of_infinite [Infinite α] : Nat.card α = 0 := mk_toNat_of_infinite +lemma cast_card [Finite α] : (Nat.card α : Cardinal) = Cardinal.mk α := by + rw [Nat.card, Cardinal.cast_toNat_of_lt_aleph0] + exact Cardinal.lt_aleph0_of_finite _ + lemma _root_.Set.Infinite.card_eq_zero {s : Set α} (hs : s.Infinite) : Nat.card s = 0 := @card_eq_zero_of_infinite _ hs.to_subtype From ab4ab8fae6e41668981caf6d986dd5eb1b577529 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Martin=20Dvo=C5=99=C3=A1k?= Date: Thu, 19 Dec 2024 13:34:02 +0000 Subject: [PATCH 784/829] docs: s/use/used (#20065) --- Mathlib/Data/Fin/Tuple/Reflection.lean | 8 ++++---- Mathlib/Data/Matrix/Reflection.lean | 4 ++-- 2 files changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Data/Fin/Tuple/Reflection.lean b/Mathlib/Data/Fin/Tuple/Reflection.lean index d0d93cedb5255..81fef2521724a 100644 --- a/Mathlib/Data/Fin/Tuple/Reflection.lean +++ b/Mathlib/Data/Fin/Tuple/Reflection.lean @@ -54,7 +54,7 @@ example {f₁ f₂ : α → β} (a₁ a₂ : α) : seq ![f₁, f₂] ![a₁, a def map (f : α → β) {m} : (Fin m → α) → Fin m → β := seq fun _ => f -/-- This can be use to prove +/-- This can be used to prove ```lean example {f : α → β} (a₁ a₂ : α) : f ∘ ![a₁, a₂] = ![f a₁, f a₂] := (map_eq _ _).symm @@ -71,7 +71,7 @@ example {f : α → β} (a₁ a₂ : α) : f ∘ ![a₁, a₂] = ![f a₁, f a def etaExpand {m} (v : Fin m → α) : Fin m → α := map id v -/-- This can be use to prove +/-- This can be used to prove ```lean example (a : Fin 2 → α) : a = ![a 0, a 1] := (etaExpand_eq _).symm @@ -89,7 +89,7 @@ def Forall : ∀ {m} (_ : (Fin m → α) → Prop), Prop | 0, P => P ![] | _ + 1, P => ∀ x : α, Forall fun v => P (Matrix.vecCons x v) -/-- This can be use to prove +/-- This can be used to prove ```lean example (P : (Fin 2 → α) → Prop) : (∀ f, P f) ↔ ∀ a₀ a₁, P ![a₀, a₁] := (forall_iff _).symm @@ -110,7 +110,7 @@ def Exists : ∀ {m} (_ : (Fin m → α) → Prop), Prop | 0, P => P ![] | _ + 1, P => ∃ x : α, Exists fun v => P (Matrix.vecCons x v) -/-- This can be use to prove +/-- This can be used to prove ```lean example (P : (Fin 2 → α) → Prop) : (∃ f, P f) ↔ ∃ a₀ a₁, P ![a₀, a₁] := (exists_iff _).symm diff --git a/Mathlib/Data/Matrix/Reflection.lean b/Mathlib/Data/Matrix/Reflection.lean index 4789ae9f2be93..ce97688fad876 100644 --- a/Mathlib/Data/Matrix/Reflection.lean +++ b/Mathlib/Data/Matrix/Reflection.lean @@ -42,7 +42,7 @@ def Forall : ∀ {m n} (_ : Matrix (Fin m) (Fin n) α → Prop), Prop | 0, _, P => P (of ![]) | _ + 1, _, P => FinVec.Forall fun r => Forall fun A => P (of (Matrix.vecCons r A)) -/-- This can be use to prove +/-- This can be used to prove ```lean example (P : Matrix (Fin 2) (Fin 3) α → Prop) : (∀ x, P x) ↔ ∀ a b c d e f, P !![a, b, c; d, e, f] := @@ -64,7 +64,7 @@ def Exists : ∀ {m n} (_ : Matrix (Fin m) (Fin n) α → Prop), Prop | 0, _, P => P (of ![]) | _ + 1, _, P => FinVec.Exists fun r => Exists fun A => P (of (Matrix.vecCons r A)) -/-- This can be use to prove +/-- This can be used to prove ```lean example (P : Matrix (Fin 2) (Fin 3) α → Prop) : (∃ x, P x) ↔ ∃ a b c d e f, P !![a, b, c; d, e, f] := From e60d040a63499f131e07742affce7fa0f41d9517 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 19 Dec 2024 13:59:50 +0000 Subject: [PATCH 785/829] chore(Order/Disjoint): generalize a lemma (#20052) Genralize `disjoint_of_le_iff_left_eq_bot` from a `SemilatticeInf` to `PartialOrder` + `OrderBot`. --- Mathlib/Data/Finset/Basic.lean | 5 +++-- Mathlib/Order/Disjoint.lean | 8 ++++---- 2 files changed, 7 insertions(+), 6 deletions(-) diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index f35079777cde0..e7fc61cbb401c 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -92,9 +92,10 @@ theorem disjoint_or_nonempty_inter (s t : Finset α) : Disjoint s t ∨ (s ∩ t rw [← not_disjoint_iff_nonempty_inter] exact em _ +omit [DecidableEq α] in theorem disjoint_of_subset_iff_left_eq_empty (h : s ⊆ t) : - Disjoint s t ↔ s = ∅ := by - rw [disjoint_iff, inf_eq_left.mpr h, bot_eq_empty] + Disjoint s t ↔ s = ∅ := + disjoint_of_le_iff_left_eq_bot h end Lattice diff --git a/Mathlib/Order/Disjoint.lean b/Mathlib/Order/Disjoint.lean index f291e96d7e476..8291d425849a7 100644 --- a/Mathlib/Order/Disjoint.lean +++ b/Mathlib/Order/Disjoint.lean @@ -90,6 +90,10 @@ lemma Disjoint.eq_iff (hab : Disjoint a b) : a = b ↔ a = ⊥ ∧ b = ⊥ := by lemma Disjoint.ne_iff (hab : Disjoint a b) : a ≠ b ↔ a ≠ ⊥ ∨ b ≠ ⊥ := hab.eq_iff.not.trans not_and_or +theorem disjoint_of_le_iff_left_eq_bot (h : a ≤ b) : + Disjoint a b ↔ a = ⊥ := + ⟨fun hd ↦ hd.eq_bot_of_le h, fun h ↦ h ▸ disjoint_bot_left⟩ + end PartialOrderBot section PartialBoundedOrder @@ -116,10 +120,6 @@ theorem disjoint_iff_inf_le : Disjoint a b ↔ a ⊓ b ≤ ⊥ := theorem disjoint_iff : Disjoint a b ↔ a ⊓ b = ⊥ := disjoint_iff_inf_le.trans le_bot_iff -theorem disjoint_of_le_iff_left_eq_bot (h : a ≤ b) : - Disjoint a b ↔ a = ⊥ := by - simp only [disjoint_iff, inf_eq_left.mpr h] - theorem Disjoint.le_bot : Disjoint a b → a ⊓ b ≤ ⊥ := disjoint_iff_inf_le.mp From 69ec65dc81216c48154474008c059388fac15246 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Martin=20Dvo=C5=99=C3=A1k?= Date: Thu, 19 Dec 2024 14:42:43 +0000 Subject: [PATCH 786/829] docs(Mathlib/LinearAlgebra/Matrix/DotProduct): remove tag reindex (#20063) --- Mathlib/LinearAlgebra/Matrix/DotProduct.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/LinearAlgebra/Matrix/DotProduct.lean b/Mathlib/LinearAlgebra/Matrix/DotProduct.lean index 0d151dcdfbceb..b2c001e265fe6 100644 --- a/Mathlib/LinearAlgebra/Matrix/DotProduct.lean +++ b/Mathlib/LinearAlgebra/Matrix/DotProduct.lean @@ -22,7 +22,7 @@ vectors `v w : n → R` to the sum of the entrywise products `v i * w i`. ## Tags -matrix, reindex +matrix -/ From 7669cc54f7c5c7f632570bfc610874709fa9306d Mon Sep 17 00:00:00 2001 From: grunweg Date: Thu, 19 Dec 2024 15:22:53 +0000 Subject: [PATCH 787/829] doc/chore: document the fields in 1000.yaml (#20066) And remove the 'identifiers' field: the field was only kept for the upstream repository, and the field was removed there. Co-authored-by: grunweg --- docs/1000.yaml | 18 +++++++++++++----- 1 file changed, 13 insertions(+), 5 deletions(-) diff --git a/docs/1000.yaml b/docs/1000.yaml index 8126486a00af9..de096445b0516 100644 --- a/docs/1000.yaml +++ b/docs/1000.yaml @@ -2,10 +2,20 @@ # 1000+ theorems project (https://1000-plus.github.io/). # This file and associated infrastructure are work in progress. # +# This file, contains for each theorem in the project, +# - an entry, labelled by its wikidata entry, +# - |title|: a human-readable title +# - (optional) |decl|: if a theorem is formalised in mathlib, the archive or counterexamples, +# the name of the corresponding declaration +# - (optional) |decls|: like |decl|, but a list of declarations +# (if one theorem is split into multiple declarations) +# - (optional) |author|: author(s) of a formalisation +# - (optional) |date|: date of completion of the first formalisation +# - (optional) |url|: for external projects, an URL referring to the result + # Current TODOs/unresolved questions: -# - should this file include further metadata, such as repository links, author names or dates? # - add infrastructure for updating the information upstream when formalisations are added -# - perhaps add a wikidata version of the stacks tag, and generate this file automatically +# - perhaps add a wikidata version of the stacks attribute, and generate this file automatically # (can this handle formalisation of *statements* also?) # - complete the information which theorems are already formalised @@ -140,9 +150,7 @@ Q207244: Q208416: title: Independence of the continuum hypothesis - url: https://github.com/flypitch/flypitch - identifiers: - - independence_of_CH + url: https://github.com/flypitch/flypitch/blob/d72904c17fbb874f01ffe168667ba12663a7b853/src/summary.lean#L90 author: Floris van Doorn and Jesse Michael Han date: 2019-09-17 From 6c80f3aa575168cc73197e1e90da5404bb371220 Mon Sep 17 00:00:00 2001 From: smorel394 Date: Thu, 19 Dec 2024 16:52:48 +0000 Subject: [PATCH 788/829] feat(CategoryTheory/Shift): commutation with pulled back shifts (#20039) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit If we have shifts by an additive monoid `B` on categories `C` and `D`, if a functor `F : C ⥤ D` commutes with the shifts, and if we have a morphism of additive monoids `φ : A →+ B`, then `F` also commutes with the pullbacks of the shifts by `φ`. Co-authored-by: smorel394 <67864981+smorel394@users.noreply.github.com> --- Mathlib/CategoryTheory/Shift/CommShift.lean | 17 ++++++ Mathlib/CategoryTheory/Shift/Pullback.lean | 67 ++++++++++++++++++++- 2 files changed, 83 insertions(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Shift/CommShift.lean b/Mathlib/CategoryTheory/Shift/CommShift.lean index 62d33df9efcfe..6aabe3d9c6729 100644 --- a/Mathlib/CategoryTheory/Shift/CommShift.lean +++ b/Mathlib/CategoryTheory/Shift/CommShift.lean @@ -47,6 +47,18 @@ noncomputable def isoZero : shiftFunctor C (0 : A) ⋙ F ≅ F ⋙ shiftFunctor isoWhiskerRight (shiftFunctorZero C A) F ≪≫ F.leftUnitor ≪≫ F.rightUnitor.symm ≪≫ isoWhiskerLeft F (shiftFunctorZero D A).symm +/-- For any functor `F : C ⥤ D` and any `a` in `A` such that `a = 0`, +this is the obvious isomorphism `shiftFunctor C a ⋙ F ≅ F ⋙ shiftFunctor D a` deduced from the +isomorphisms `shiftFunctorZero'` on both categories `C` and `D`. -/ +@[simps!] +noncomputable def isoZero' (a : A) (ha : a = 0) : shiftFunctor C a ⋙ F ≅ F ⋙ shiftFunctor D a := + isoWhiskerRight (shiftFunctorZero' C a ha) F ≪≫ F.leftUnitor ≪≫ + F.rightUnitor.symm ≪≫ isoWhiskerLeft F (shiftFunctorZero' D a ha).symm + +@[simp] +lemma isoZero'_eq_isoZero : isoZero' F A 0 rfl = isoZero F A := by + ext; simp [isoZero', shiftFunctorZero'] + variable {F A} /-- If a functor `F : C ⥤ D` is equipped with "commutation isomorphisms" with the @@ -132,6 +144,11 @@ lemma commShiftIso_zero : F.commShiftIso (0 : A) = CommShift.isoZero F A := CommShift.zero +set_option linter.docPrime false in +lemma commShiftIso_zero' (a : A) (h : a = 0) : + F.commShiftIso a = CommShift.isoZero' F A a h := by + subst h; rw [CommShift.isoZero'_eq_isoZero, commShiftIso_zero] + variable {A} lemma commShiftIso_add (a b : A) : diff --git a/Mathlib/CategoryTheory/Shift/Pullback.lean b/Mathlib/CategoryTheory/Shift/Pullback.lean index aebfda6c350d3..0bce51f3ba6af 100644 --- a/Mathlib/CategoryTheory/Shift/Pullback.lean +++ b/Mathlib/CategoryTheory/Shift/Pullback.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Joël Riou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ -import Mathlib.CategoryTheory.Shift.Basic +import Mathlib.CategoryTheory.Shift.CommShift import Mathlib.CategoryTheory.Preadditive.AdditiveFunctor /-! @@ -74,6 +74,21 @@ lemma pullbackShiftFunctorZero_hom_app : pullbackShiftFunctorZero_inv_app, assoc, Iso.inv_hom_id_app_assoc, Iso.inv_hom_id_app] rfl +lemma pullbackShiftFunctorZero'_inv_app : + (shiftFunctorZero _ A).inv.app X = (shiftFunctorZero' C (φ 0) (by rw [map_zero])).inv.app X ≫ + (pullbackShiftIso C φ 0 (φ 0) rfl).inv.app X := by + rw [pullbackShiftFunctorZero_inv_app] + simp only [Functor.id_obj, pullbackShiftIso, eqToIso.inv, eqToHom_app, shiftFunctorZero', + Iso.trans_inv, NatTrans.comp_app, eqToIso_refl, Iso.refl_inv, NatTrans.id_app, assoc] + erw [comp_id] + +lemma pullbackShiftFunctorZero'_hom_app : + (shiftFunctorZero _ A).hom.app X = (pullbackShiftIso C φ 0 (φ 0) rfl).hom.app X ≫ + (shiftFunctorZero' C (φ 0) (by rw [map_zero])).hom.app X := by + rw [← cancel_epi ((shiftFunctorZero _ A).inv.app X), Iso.inv_hom_id_app, + pullbackShiftFunctorZero'_inv_app, assoc, Iso.inv_hom_id_app_assoc, Iso.inv_hom_id_app] + rfl + lemma pullbackShiftFunctorAdd'_inv_app : (shiftFunctorAdd' _ a₁ a₂ a₃ h).inv.app X = (shiftFunctor (PullbackShift C φ) a₂).map ((pullbackShiftIso C φ a₁ b₁ h₁).hom.app X) ≫ @@ -103,4 +118,54 @@ lemma pullbackShiftFunctorAdd'_hom_app : ← Functor.map_comp, Iso.hom_inv_id_app, Functor.map_id] rfl +namespace Functor + +variable {D : Type*} [Category D] [HasShift D B] (F : C ⥤ D) [F.CommShift B] + +/-- If `F : C ⥤ D` commutes with the shifts on `C` and `D`, then it also commutes with +their pullbacks by an additive map. +-/ +noncomputable def commShiftPullback : + F.CommShift A (C := PullbackShift C φ) (D := PullbackShift D φ) where + iso a := isoWhiskerRight (pullbackShiftIso C φ a (φ a) rfl) F ≪≫ + F.commShiftIso (φ a) ≪≫ isoWhiskerLeft _ (pullbackShiftIso D φ a (φ a) rfl).symm + zero := by + ext + dsimp + simp only [F.commShiftIso_zero' (A := B) (φ 0) (by rw [map_zero]), CommShift.isoZero'_hom_app, + assoc, CommShift.isoZero_hom_app, pullbackShiftFunctorZero'_hom_app, map_comp, + pullbackShiftFunctorZero'_inv_app] + dsimp + rfl + add a b := by + ext + dsimp + simp only [CommShift.isoAdd_hom_app, map_comp, assoc] + dsimp + rw [F.commShiftIso_add' (a := φ a) (b := φ b) (by rw [φ.map_add]), + ← shiftFunctorAdd'_eq_shiftFunctorAdd, ← shiftFunctorAdd'_eq_shiftFunctorAdd, + pullbackShiftFunctorAdd'_hom_app φ _ a b (a + b) rfl (φ a) (φ b) (φ (a + b)) rfl rfl rfl, + pullbackShiftFunctorAdd'_inv_app φ _ a b (a + b) rfl (φ a) (φ b) (φ (a + b)) rfl rfl rfl] + dsimp + simp only [CommShift.isoAdd'_hom_app, assoc, map_comp, NatTrans.naturality_assoc, + Iso.inv_hom_id_app_assoc] + slice_rhs 9 10 => rw [← map_comp, Iso.inv_hom_id_app, map_id] + erw [id_comp] + slice_rhs 6 7 => erw [← (CommShift.iso (φ b)).hom.naturality] + slice_rhs 4 5 => rw [← map_comp, (pullbackShiftIso C φ b (φ b) rfl).hom.naturality, map_comp] + simp only [comp_obj, Functor.comp_map, assoc] + slice_rhs 3 4 => rw [← map_comp, Iso.inv_hom_id_app, map_id] + slice_rhs 4 5 => rw [← map_comp]; erw [← map_comp]; rw [Iso.inv_hom_id_app, map_id, map_id] + rw [id_comp, id_comp, assoc, assoc]; rfl + +lemma commShiftPullback_iso_eq (a : A) (b : B) (h : b = φ a) : + letI : F.CommShift (C := PullbackShift C φ) (D := PullbackShift D φ) A := F.commShiftPullback φ + F.commShiftIso a (C := PullbackShift C φ) (D := PullbackShift D φ) = + isoWhiskerRight (pullbackShiftIso C φ a b h) F ≪≫ (F.commShiftIso b) ≪≫ + isoWhiskerLeft F (pullbackShiftIso D φ a b h).symm := by + obtain rfl : b = φ a := h + rfl + +end Functor + end CategoryTheory From 51104eb3a7b79723dcd6d82ac5ec40c26257f623 Mon Sep 17 00:00:00 2001 From: Arend Mellendijk Date: Thu, 19 Dec 2024 17:17:06 +0000 Subject: [PATCH 789/829] feat(NumberTheory/ArithmeticFunction): Miscellaneous lemmas about multiplicative arithmetic functions (#20025) These are three small lemmas that were helpful while developing the Selberg sieve. --- Mathlib/NumberTheory/ArithmeticFunction.lean | 24 ++++++++++++++++++++ 1 file changed, 24 insertions(+) diff --git a/Mathlib/NumberTheory/ArithmeticFunction.lean b/Mathlib/NumberTheory/ArithmeticFunction.lean index 8b8038b188ab0..900057426c931 100644 --- a/Mathlib/NumberTheory/ArithmeticFunction.lean +++ b/Mathlib/NumberTheory/ArithmeticFunction.lean @@ -590,6 +590,12 @@ theorem map_prod_of_subset_primeFactors [CommSemiring R] {f : ArithmeticFunction f (∏ a ∈ t, a) = ∏ a ∈ t, f a := map_prod_of_prime h_mult t fun _ a => prime_of_mem_primeFactors (ht a) +theorem map_div_of_coprime [CommGroupWithZero R] {f : ArithmeticFunction R} + (hf : IsMultiplicative f) {l d : ℕ} (hdl : d ∣ l) (hl : (l/d).Coprime d) (hd : f d ≠ 0) : + f (l / d) = f l / f d := by + apply (div_eq_of_eq_mul hd ..).symm + rw [← hf.right hl, Nat.div_mul_cancel hdl] + @[arith_mult] theorem natCast {f : ArithmeticFunction ℕ} [Semiring R] (h : f.IsMultiplicative) : IsMultiplicative (f : ArithmeticFunction R) := @@ -771,6 +777,24 @@ theorem lcm_apply_mul_gcd_apply [CommMonoidWithZero R] {f : ArithmeticFunction R apply Finset.inter_subset_union · simp [factorization_lcm hx hy] +theorem map_gcd [CommGroupWithZero R] {f : ArithmeticFunction R} + (hf : f.IsMultiplicative) {x y : ℕ} (hf_lcm : f (x.lcm y) ≠ 0) : + f (x.gcd y) = f x * f y / f (x.lcm y) := by + rw [←hf.lcm_apply_mul_gcd_apply, mul_div_cancel_left₀ _ hf_lcm] + +theorem map_lcm [CommGroupWithZero R] {f : ArithmeticFunction R} + (hf : f.IsMultiplicative) {x y : ℕ} (hf_gcd : f (x.gcd y) ≠ 0) : + f (x.lcm y) = f x * f y / f (x.gcd y) := by + rw [←hf.lcm_apply_mul_gcd_apply, mul_div_cancel_right₀ _ hf_gcd] + +theorem eq_zero_of_squarefree_of_dvd_eq_zero [CommMonoidWithZero R] {f : ArithmeticFunction R} + (hf : IsMultiplicative f) {m n : ℕ} (hn : Squarefree n) (hmn : m ∣ n) + (h_zero : f m = 0) : + f n = 0 := by + rcases hmn with ⟨k, rfl⟩ + simp only [MulZeroClass.zero_mul, eq_self_iff_true, hf.map_mul_of_coprime + (coprime_of_squarefree_mul hn), h_zero] + end IsMultiplicative section SpecialFunctions From 49ac6bf04dc187ffc0a056421c603fcefce6991a Mon Sep 17 00:00:00 2001 From: smorel394 <67864981+smorel394@users.noreply.github.com> Date: Thu, 19 Dec 2024 17:27:28 +0000 Subject: [PATCH 790/829] feat(CategoryTheory/Shift): commutation with shifts and adjunctions (#20033) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Given categories `C` and `D` that have shifts by an additive group `A`, functors `F : C ⥤ D` and `G : C ⥤ D`, an adjunction `F ⊣ G` and a `CommShift` structure on `F`, construct a `CommShift` structure on `G`. As an easy application, if `E : C ≌ D` is an equivalence and `E.functor` has a `CommShift` structure, we get a `CommShift` structure on `E.inverse`. The `CommShift` structure on `G` must be compatible with the one on `F` in the following sense (cf. `Adjunction.CommShift`): for every `a` in `A`, the natural transformation `adj.unit : 𝟭 C ⟶ G ⋙ F` commutes with the isomorphism `shiftFunctor C A ⋙ G ⋙ F ≅ G ⋙ F ⋙ shiftFunctor C A` induces by `F.commShiftIso a` and `G.commShiftIso a`. We actually require a similar condition for `adj.counit`, but it follows from the one for `adj.unit`. In order to simplify the construction of the `CommShift` structure on `G`, we first introduce the compatibility condition on `adj.unit` for a fixed `a` in `A` and for isomorphisms `e₁ : shiftFunctor C a ⋙ F ≅ F ⋙ shiftFunctor D a` and `e₂ : shiftFunctor D a ⋙ G ≅ G ⋙ shiftFunctor C a`. We then prove that: - If `e₁` and `e₂` satusfy this condition, then `e₁` uniquely determines `e₂` and vice versa. - If `a = 0`, the isomorphisms `Functor.CommShift.isoZero F` and `Functor.CommShift.isoZero G` satisfy the condition. - The condition is stable by addition on `A`, if we use `Functor.CommShift.isoAdd` to deduce commutation isomorphism for `a + b` from such isomorphism from `a` and `b`. - Given commutation isomorphisms for `F`, our candidate commutation isomorphisms for `G`, constructed in `Adjunction.RightAdjointCommShift.iso`, satisfy the compatibility condition. Once we have established all this, the compatibility of the commutation isomorphism for `F` expressed in `CommShift.zero` and `CommShift.add` immediately implies the similar statements for the commutation isomorphisms for `G`. TODO: Construct a `CommShift` structure on `F` from a `CommShift` structure on `G`, using opposite categories. Co-authored-by: morel --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Shift/Adjunction.lean | 436 +++++++++++++++++++ 2 files changed, 437 insertions(+) create mode 100644 Mathlib/CategoryTheory/Shift/Adjunction.lean diff --git a/Mathlib.lean b/Mathlib.lean index ac8fe9eacb347..fc9b7328725b2 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2060,6 +2060,7 @@ import Mathlib.CategoryTheory.Quotient import Mathlib.CategoryTheory.Quotient.Linear import Mathlib.CategoryTheory.Quotient.Preadditive import Mathlib.CategoryTheory.Retract +import Mathlib.CategoryTheory.Shift.Adjunction import Mathlib.CategoryTheory.Shift.Basic import Mathlib.CategoryTheory.Shift.CommShift import Mathlib.CategoryTheory.Shift.Induced diff --git a/Mathlib/CategoryTheory/Shift/Adjunction.lean b/Mathlib/CategoryTheory/Shift/Adjunction.lean new file mode 100644 index 0000000000000..0cbde40a773ff --- /dev/null +++ b/Mathlib/CategoryTheory/Shift/Adjunction.lean @@ -0,0 +1,436 @@ +/- +Copyright (c) 2024 Sophie Morel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sophie Morel, Joël Riou +-/ +import Mathlib.CategoryTheory.Shift.CommShift +import Mathlib.CategoryTheory.Adjunction.Mates + +/-! +# Adjoints commute with shifts + +Given categories `C` and `D` that have shifts by an additive group `A`, functors `F : C ⥤ D` +and `G : C ⥤ D`, an adjunction `F ⊣ G` and a `CommShift` structure on `F`, this file constructs +a `CommShift` structure on `G`. As an easy application, if `E : C ≌ D` is an equivalence and +`E.functor` has a `CommShift` structure, we get a `CommShift` structure on `E.inverse`. + +The `CommShift` structure on `G` must be compatible with the one on `F` in the following sense +(cf. `Adjunction.CommShift`): +for every `a` in `A`, the natural transformation `adj.unit : 𝟭 C ⟶ G ⋙ F` commutes with +the isomorphism `shiftFunctor C A ⋙ G ⋙ F ≅ G ⋙ F ⋙ shiftFunctor C A` induces by +`F.commShiftIso a` and `G.commShiftIso a`. We actually require a similar condition for +`adj.counit`, but it follows from the one for `adj.unit`. + +In order to simplify the construction of the `CommShift` structure on `G`, we first introduce +the compatibility condition on `adj.unit` for a fixed `a` in `A` and for isomorphisms +`e₁ : shiftFunctor C a ⋙ F ≅ F ⋙ shiftFunctor D a` and +`e₂ : shiftFunctor D a ⋙ G ≅ G ⋙ shiftFunctor C a`. We then prove that: +- If `e₁` and `e₂` satusfy this condition, then `e₁` uniquely determines `e₂` and vice versa. +- If `a = 0`, the isomorphisms `Functor.CommShift.isoZero F` and `Functor.CommShift.isoZero G` +satisfy the condition. +- The condition is stable by addition on `A`, if we use `Functor.CommShift.isoAdd` to deduce +commutation isomorphism for `a + b` from such isomorphism from `a` and `b`. +- Given commutation isomorphisms for `F`, our candidate commutation isomorphisms for `G`, +constructed in `Adjunction.RightAdjointCommShift.iso`, satisfy the compatibility condition. + +Once we have established all this, the compatibility of the commutation isomorphism for +`F` expressed in `CommShift.zero` and `CommShift.add` immediately implies the similar +statements for the commutation isomorphisms for `G`. + +TODO: Construct a `CommShift` structure on `F` from a `CommShift` structure on `G`, using +opposite categories. + +-/ + +namespace CategoryTheory + +open Category + +namespace Adjunction + +variable {C D : Type*} [Category C] [Category D] + {F : C ⥤ D} {G : D ⥤ C} (adj : F ⊣ G) {A : Type*} [AddMonoid A] [HasShift C A] [HasShift D A] + +namespace CommShift + +variable {a b : A} (e₁ : shiftFunctor C a ⋙ F ≅ F ⋙ shiftFunctor D a) + (e₁' : shiftFunctor C a ⋙ F ≅ F ⋙ shiftFunctor D a) + (f₁ : shiftFunctor C b ⋙ F ≅ F ⋙ shiftFunctor D b) + (e₂ : shiftFunctor D a ⋙ G ≅ G ⋙ shiftFunctor C a) + (e₂' : shiftFunctor D a ⋙ G ≅ G ⋙ shiftFunctor C a) + (f₂ : shiftFunctor D b ⋙ G ≅ G ⋙ shiftFunctor C b) + +/-- Given an adjunction `adj : F ⊣ G`, `a` in `A` and commutation isomorphisms +`e₁ : shiftFunctor C a ⋙ F ≅ F ⋙ shiftFunctor D a` and +`e₂ : shiftFunctor D a ⋙ G ≅ G ⋙ shiftFunctor C a`, this expresses the compatibility of +`e₁` and `e₂` with the unit of the adjunction `adj`. +-/ +abbrev CompatibilityUnit := + ∀ (X : C), (adj.unit.app X)⟦a⟧' = adj.unit.app (X⟦a⟧) ≫ G.map (e₁.hom.app X) ≫ e₂.hom.app _ + +/-- Given an adjunction `adj : F ⊣ G`, `a` in `A` and commutation isomorphisms +`e₁ : shiftFunctor C a ⋙ F ≅ F ⋙ shiftFunctor D a` and +`e₂ : shiftFunctor D a ⋙ G ≅ G ⋙ shiftFunctor C a`, this expresses the compatibility of +`e₁` and `e₂` with the counit of the adjunction `adj`. +-/ +abbrev CompatibilityCounit := + ∀ (Y : D), adj.counit.app (Y⟦a⟧) = F.map (e₂.hom.app Y) ≫ e₁.hom.app _ ≫ (adj.counit.app Y)⟦a⟧' + +/-- Given an adjunction `adj : F ⊣ G`, `a` in `A` and commutation isomorphisms +`e₁ : shiftFunctor C a ⋙ F ≅ F ⋙ shiftFunctor D a` and +`e₂ : shiftFunctor D a ⋙ G ≅ G ⋙ shiftFunctor C a`, compatibility of `e₁` and `e₂` with the +unit of the adjunction `adj` implies compatibility with the counit of `adj`. +-/ +lemma compatibilityCounit_of_compatibilityUnit (h : CompatibilityUnit adj e₁ e₂) : + CompatibilityCounit adj e₁ e₂ := by + intro Y + have eq := h (G.obj Y) + simp only [← cancel_mono (e₂.inv.app _ ≫ G.map (e₁.inv.app _)), + assoc, Iso.hom_inv_id_app_assoc, comp_id, ← Functor.map_comp, + Iso.hom_inv_id_app, Functor.comp_obj, Functor.map_id] at eq + apply (adj.homEquiv _ _).injective + dsimp + rw [adj.homEquiv_unit, adj.homEquiv_unit, G.map_comp, adj.unit_naturality_assoc, ← eq] + simp only [assoc, ← Functor.map_comp, Iso.inv_hom_id_app_assoc] + erw [← e₂.inv.naturality] + dsimp + simp only [right_triangle_components, ← Functor.map_comp_assoc, Functor.map_id, id_comp, + Iso.hom_inv_id_app, Functor.comp_obj] + +/-- Given an adjunction `adj : F ⊣ G`, `a` in `A` and commutation isomorphisms +`e₁ : shiftFunctor C a ⋙ F ≅ F ⋙ shiftFunctor D a` and +`e₂ : shiftFunctor D a ⋙ G ≅ G ⋙ shiftFunctor C a`, if `e₁` and `e₂` are compatible with the +unit of the adjunction `adj`, then we get a formula for `e₂.inv` in terms of `e₁`. +-/ +lemma compatibilityUnit_right (h : CompatibilityUnit adj e₁ e₂) (Y : D) : + e₂.inv.app Y = adj.unit.app _ ≫ G.map (e₁.hom.app _) ≫ G.map ((adj.counit.app _)⟦a⟧') := by + have := h (G.obj Y) + rw [← cancel_mono (e₂.inv.app _), assoc, assoc, Iso.hom_inv_id_app] at this + erw [comp_id] at this + rw [← assoc, ← this, assoc]; erw [← e₂.inv.naturality] + rw [← cancel_mono (e₂.hom.app _)] + simp only [Functor.comp_obj, Iso.inv_hom_id_app, Functor.id_obj, Functor.comp_map, assoc, comp_id, + ← (shiftFunctor C a).map_comp, right_triangle_components, Functor.map_id] + +/-- Given an adjunction `adj : F ⊣ G`, `a` in `A` and commutation isomorphisms +`e₁ : shiftFunctor C a ⋙ F ≅ F ⋙ shiftFunctor D a` and +`e₂ : shiftFunctor D a ⋙ G ≅ G ⋙ shiftFunctor C a`, if `e₁` and `e₂` are compatible with the +counit of the adjunction `adj`, then we get a formula for `e₁.hom` in terms of `e₂`. +-/ +lemma compatibilityCounit_left (h : CompatibilityCounit adj e₁ e₂) (X : C) : + e₁.hom.app X = F.map ((adj.unit.app X)⟦a⟧') ≫ F.map (e₂.inv.app _) ≫ adj.counit.app _ := by + have := h (F.obj X) + rw [← cancel_epi (F.map (e₂.inv.app _)), ← assoc, ← F.map_comp, Iso.inv_hom_id_app, F.map_id, + id_comp] at this + rw [this] + erw [e₁.hom.naturality_assoc] + rw [Functor.comp_map, ← Functor.map_comp, left_triangle_components] + simp only [Functor.comp_obj, Functor.id_obj, Functor.map_id, comp_id] + +/-- Given an adjunction `adj : F ⊣ G`, `a` in `A` and commutation isomorphisms +`e₁ : shiftFunctor C a ⋙ F ≅ F ⋙ shiftFunctor D a` and +`e₂ : shiftFunctor D a ⋙ G ≅ G ⋙ shiftFunctor C a`, if `e₁` and `e₂` are compatible with the +unit of the adjunction `adj`, then `e₁` uniquely determines `e₂`. +-/ +lemma compatibilityUnit_unique_right (h : CompatibilityUnit adj e₁ e₂) + (h' : CompatibilityUnit adj e₁ e₂') : e₂ = e₂' := by + rw [← Iso.symm_eq_iff] + ext + rw [Iso.symm_hom, Iso.symm_hom, compatibilityUnit_right adj e₁ e₂ h, + compatibilityUnit_right adj e₁ e₂' h'] + +/-- Given an adjunction `adj : F ⊣ G`, `a` in `A` and commutation isomorphisms +`e₁ : shiftFunctor C a ⋙ F ≅ F ⋙ shiftFunctor D a` and +`e₂ : shiftFunctor D a ⋙ G ≅ G ⋙ shiftFunctor C a`, if `e₁` and `e₂` are compatible with the +counit of the adjunction `adj`, then `e₂` uniquely determines `e₁`. +-/ +lemma compatibilityCounit_unique_left (h : CompatibilityCounit adj e₁ e₂) + (h' : CompatibilityCounit adj e₁' e₂) : e₁ = e₁' := by + ext + rw [compatibilityCounit_left adj e₁ e₂ h, compatibilityCounit_left adj e₁' e₂ h'] + +/-- +The isomorphisms `Functor.CommShift.isoZero F` and `Functor.CommShift.isoZero G` are +compatible with the unit of an adjunction `F ⊣ G`. +-/ +lemma compatibilityUnit_isoZero : CompatibilityUnit adj (Functor.CommShift.isoZero F A) + (Functor.CommShift.isoZero G A) := by + intro + simp only [Functor.id_obj, Functor.comp_obj, Functor.CommShift.isoZero_hom_app, + Functor.map_comp, assoc, unit_naturality_assoc, + ← cancel_mono ((shiftFunctorZero C A).hom.app _), ← G.map_comp_assoc, Iso.inv_hom_id_app, + Functor.id_obj, Functor.map_id, id_comp, NatTrans.naturality, Functor.id_map, assoc, comp_id] + +/-- Given an adjunction `adj : F ⊣ G`, `a, b` in `A` and commutation isomorphisms +between shifts by `a` (resp. `b`) and `F` and `G`, if these commutation isomorphisms are +compatible with the unit of `adj`, then so are the commutation isomorphisms between shifts +by `a + b` and `F` and `G` constructed by `Functor.CommShift.isoAdd`. +-/ +lemma compatibilityUnit_isoAdd (h : CompatibilityUnit adj e₁ e₂) + (h' : CompatibilityUnit adj f₁ f₂) : + CompatibilityUnit adj (Functor.CommShift.isoAdd e₁ f₁) (Functor.CommShift.isoAdd e₂ f₂) := by + intro X + have := h' (X⟦a⟧) + simp only [← cancel_mono (f₂.inv.app _), assoc, Iso.hom_inv_id_app, + Functor.id_obj, Functor.comp_obj, comp_id] at this + simp only [Functor.id_obj, Functor.comp_obj, Functor.CommShift.isoAdd_hom_app, + Functor.map_comp, assoc, unit_naturality_assoc] + slice_rhs 5 6 => rw [← G.map_comp, Iso.inv_hom_id_app] + simp only [Functor.comp_obj, Functor.map_id, id_comp, assoc] + erw [f₂.hom.naturality_assoc] + rw [← reassoc_of% this, ← cancel_mono ((shiftFunctorAdd C a b).hom.app _), + assoc, assoc, assoc, assoc, assoc, assoc, Iso.inv_hom_id_app_assoc, Iso.inv_hom_id_app] + dsimp + rw [← (shiftFunctor C b).map_comp_assoc, ← (shiftFunctor C b).map_comp_assoc, + assoc, ← h X, NatTrans.naturality] + dsimp + rw [comp_id] + +end CommShift + +variable (A) [F.CommShift A] [G.CommShift A] + +/-- +The property for `CommShift` structures on `F` and `G` to be compatible with an +adjunction `F ⊣ G`. +-/ +class CommShift : Prop where + commShift_unit : NatTrans.CommShift adj.unit A := by infer_instance + commShift_counit : NatTrans.CommShift adj.counit A := by infer_instance + +namespace CommShift + +attribute [instance] commShift_unit commShift_counit + +/-- Constructor for `Adjunction.CommShift`. -/ +lemma mk' (h : NatTrans.CommShift adj.unit A) : + adj.CommShift A where + commShift_counit := ⟨fun a ↦ by + ext + simp only [Functor.comp_obj, Functor.id_obj, NatTrans.comp_app, + Functor.commShiftIso_comp_hom_app, whiskerRight_app, assoc, whiskerLeft_app, + Functor.commShiftIso_id_hom_app, comp_id] + refine (compatibilityCounit_of_compatibilityUnit adj _ _ (fun X ↦ ?_) _).symm + simpa only [NatTrans.comp_app, + Functor.commShiftIso_id_hom_app, whiskerRight_app, id_comp, + Functor.commShiftIso_comp_hom_app] using congr_app (h.comm' a) X⟩ + +end CommShift + +variable {A} + +@[reassoc] +lemma shift_unit_app [adj.CommShift A] (a : A) (X : C) : + (adj.unit.app X)⟦a⟧' = + adj.unit.app (X⟦a⟧) ≫ + G.map ((F.commShiftIso a).hom.app X) ≫ + (G.commShiftIso a).hom.app (F.obj X) := by + simpa [Functor.commShiftIso_comp_hom_app] using NatTrans.CommShift.comm_app adj.unit a X + +@[reassoc] +lemma shift_counit_app [adj.CommShift A] (a : A) (Y : D) : + (adj.counit.app Y)⟦a⟧' = + (F.commShiftIso a).inv.app (G.obj Y) ≫ F.map ((G.commShiftIso a).inv.app Y) ≫ + adj.counit.app (Y⟦a⟧) := by + have eq := NatTrans.CommShift.comm_app adj.counit a Y + simp only [Functor.comp_obj, Functor.id_obj, Functor.commShiftIso_comp_hom_app, assoc, + Functor.commShiftIso_id_hom_app, comp_id] at eq + simp only [← eq, Functor.comp_obj, Functor.id_obj, ← F.map_comp_assoc, Iso.inv_hom_id_app, + F.map_id, id_comp, Iso.inv_hom_id_app_assoc] + +end Adjunction + +namespace Adjunction + +variable {C D : Type*} [Category C] [Category D] + {F : C ⥤ D} {G : D ⥤ C} (adj : F ⊣ G) {A : Type*} [AddGroup A] [HasShift C A] [HasShift D A] + +namespace RightAdjointCommShift + +variable (a b : A) (h : b + a = 0) [F.CommShift A] + +/-- Auxiliary definition for `iso`. -/ +noncomputable def iso' : shiftFunctor D a ⋙ G ≅ G ⋙ shiftFunctor C a := + (conjugateIsoEquiv (Adjunction.comp adj (shiftEquiv' D b a h).toAdjunction) + (Adjunction.comp (shiftEquiv' C b a h).toAdjunction adj)).toFun (F.commShiftIso b) + +/-- +Given an adjunction `F ⊣ G` and a `CommShift` structure on `F`, these are the candidate +`CommShift.iso a` isomorphisms for a compatible `CommShift` structure on `G`. +-/ +noncomputable def iso : shiftFunctor D a ⋙ G ≅ G ⋙ shiftFunctor C a := + iso' adj _ _ (neg_add_cancel a) + +@[reassoc] +lemma iso_hom_app (X : D) : + (iso adj a).hom.app X = + (shiftFunctorCompIsoId C b a h).inv.app (G.obj ((shiftFunctor D a).obj X)) ≫ + (adj.unit.app ((shiftFunctor C b).obj (G.obj ((shiftFunctor D a).obj X))))⟦a⟧' ≫ + (G.map ((F.commShiftIso b).hom.app (G.obj ((shiftFunctor D a).obj X))))⟦a⟧' ≫ + (G.map ((shiftFunctor D b).map (adj.counit.app ((shiftFunctor D a).obj X))))⟦a⟧' ≫ + (G.map ((shiftFunctorCompIsoId D a b + (by rw [← add_left_inj a, add_assoc, h, zero_add, add_zero])).hom.app X))⟦a⟧' := by + obtain rfl : b = -a := by rw [← add_left_inj a, h, neg_add_cancel] + simp [iso, iso'] + rfl + +@[reassoc] +lemma iso_inv_app (Y : D) : + (iso adj a).inv.app Y = + adj.unit.app ((shiftFunctor C a).obj (G.obj Y)) ≫ + G.map ((shiftFunctorCompIsoId D b a h).inv.app + (F.obj ((shiftFunctor C a).obj (G.obj Y)))) ≫ + G.map ((shiftFunctor D a).map ((shiftFunctor D b).map + ((F.commShiftIso a).hom.app (G.obj Y)))) ≫ + G.map ((shiftFunctor D a).map ((shiftFunctorCompIsoId D a b + (by rw [eq_neg_of_add_eq_zero_left h, add_neg_cancel])).hom.app + (F.obj (G.obj Y)))) ≫ + G.map ((shiftFunctor D a).map (adj.counit.app Y)) := by + obtain rfl : b = -a := by rw [← add_left_inj a, h, neg_add_cancel] + simp only [Functor.comp_obj, iso, iso', shiftEquiv', Equiv.toFun_as_coe, + conjugateIsoEquiv_apply_inv, conjugateEquiv_apply_app, comp_unit_app, Functor.id_obj, + Equivalence.toAdjunction_unit, Equivalence.Equivalence_mk'_unit, Iso.symm_hom, Functor.comp_map, + comp_counit_app, Equivalence.toAdjunction_counit, Equivalence.Equivalence_mk'_counit, + Functor.map_shiftFunctorCompIsoId_hom_app, assoc, Functor.map_comp] + slice_lhs 3 4 => rw [← Functor.map_comp, ← Functor.map_comp, Iso.inv_hom_id_app] + simp only [Functor.comp_obj, Functor.map_id, id_comp, assoc] + +/-- +The commutation isomorphisms of `Adjunction.RightAdjointCommShift.iso` are compatible with +the unit of the adjunction. +-/ +lemma compatibilityUnit_iso (a : A) : + CommShift.CompatibilityUnit adj (F.commShiftIso a) (iso adj a) := by + intro + rw [← cancel_mono ((RightAdjointCommShift.iso adj a).inv.app _), assoc, assoc, + Iso.hom_inv_id_app, RightAdjointCommShift.iso_inv_app adj _ _ (neg_add_cancel a)] + apply (adj.homEquiv _ _).symm.injective + dsimp + simp only [comp_id, homEquiv_counit, Functor.map_comp, assoc, counit_naturality, + counit_naturality_assoc, left_triangle_components_assoc] + erw [← NatTrans.naturality_assoc] + dsimp + rw [shift_shiftFunctorCompIsoId_hom_app, Iso.inv_hom_id_app_assoc, + Functor.commShiftIso_hom_naturality_assoc, ← Functor.map_comp, + left_triangle_components, Functor.map_id, comp_id] + +end RightAdjointCommShift + +variable (A) + +open RightAdjointCommShift in +/-- +Given an adjunction `F ⊣ G` and a `CommShift` structure on `F`, this constructs +the unique compatible `CommShift` structure on `G`. +-/ +@[simps] +noncomputable def rightAdjointCommShift [F.CommShift A] : G.CommShift A where + iso a := iso adj a + zero := by + refine CommShift.compatibilityUnit_unique_right adj (F.commShiftIso 0) _ _ + (compatibilityUnit_iso adj 0) ?_ + rw [F.commShiftIso_zero] + exact CommShift.compatibilityUnit_isoZero adj + add a b := by + refine CommShift.compatibilityUnit_unique_right adj (F.commShiftIso (a + b)) _ _ + (compatibilityUnit_iso adj (a + b)) ?_ + rw [F.commShiftIso_add] + exact CommShift.compatibilityUnit_isoAdd adj _ _ _ _ + (compatibilityUnit_iso adj a) (compatibilityUnit_iso adj b) + +lemma commShift_of_leftAdjoint [F.CommShift A] : + letI := adj.rightAdjointCommShift A + adj.CommShift A := by + letI := adj.rightAdjointCommShift A + refine CommShift.mk' _ _ ⟨fun a ↦ ?_⟩ + ext X + dsimp + simpa only [Functor.commShiftIso_id_hom_app, Functor.comp_obj, Functor.id_obj, id_comp, + Functor.commShiftIso_comp_hom_app] using RightAdjointCommShift.compatibilityUnit_iso adj a X + +end Adjunction + +namespace Equivalence + +variable {C D : Type*} [Category C] [Category D] (E : C ≌ D) + +section + +variable (A : Type*) [AddMonoid A] [HasShift C A] [HasShift D A] + +/-- +If `E : C ≌ D` is an equivalence, this expresses the compatibility of `CommShift` +structures on `E.functor` and `E.inverse`. +-/ +abbrev CommShift [E.functor.CommShift A] [E.inverse.CommShift A] : Prop := + E.toAdjunction.CommShift A + +namespace CommShift + +variable [E.functor.CommShift A] [E.inverse.CommShift A] + +instance [E.CommShift A] : NatTrans.CommShift E.unitIso.hom A := + inferInstanceAs (NatTrans.CommShift E.toAdjunction.unit A) + +instance [E.CommShift A] : NatTrans.CommShift E.counitIso.hom A := + inferInstanceAs (NatTrans.CommShift E.toAdjunction.counit A) + +instance [h : E.functor.CommShift A] : E.symm.inverse.CommShift A := h +instance [h : E.inverse.CommShift A] : E.symm.functor.CommShift A := h + +/-- Constructor for `Equivalence.CommShift`. -/ +lemma mk' (h : NatTrans.CommShift E.unitIso.hom A) : + E.CommShift A where + commShift_unit := h + commShift_counit := (Adjunction.CommShift.mk' E.toAdjunction A h).commShift_counit + +/-- +If `E : C ≌ D` is an equivalence and we have compatible `CommShift` structures on `E.functor` +and `E.inverse`, then we also have compatible `CommShift` structures on `E.symm.functor` +and `E.symm.inverse`. +-/ +instance [E.CommShift A] : E.symm.CommShift A := + mk' E.symm A (inferInstanceAs (NatTrans.CommShift E.counitIso.inv A)) + +/-- Constructor for `Equivalence.CommShift`. -/ +lemma mk'' (h : NatTrans.CommShift E.counitIso.hom A) : + E.CommShift A := + have := mk' E.symm A (inferInstanceAs (NatTrans.CommShift E.counitIso.inv A)) + inferInstanceAs (E.symm.symm.CommShift A) + +end CommShift + +end + +variable (A : Type*) [AddGroup A] [HasShift C A] [HasShift D A] + +/-- +If `E : C ≌ D` is an equivalence and we have a `CommShift` structure on `E.functor`, +this constructs the unique compatible `CommShift` structure on `E.inverse`. +-/ +noncomputable def commShiftInverse [E.functor.CommShift A] : E.inverse.CommShift A := + E.toAdjunction.rightAdjointCommShift A + +lemma commShift_of_functor [E.functor.CommShift A] : + letI := E.commShiftInverse A + E.CommShift A := by + letI := E.commShiftInverse A + exact CommShift.mk' _ _ (E.toAdjunction.commShift_of_leftAdjoint A).commShift_unit + +/-- +If `E : C ≌ D` is an equivalence and we have a `CommShift` structure on `E.inverse`, +this constructs the unique compatible `CommShift` structure on `E.functor`. +-/ +noncomputable def commShiftFunctor [E.inverse.CommShift A] : E.functor.CommShift A := + E.symm.toAdjunction.rightAdjointCommShift A + +lemma commShift_of_inverse [E.inverse.CommShift A] : + letI := E.commShiftFunctor A + E.CommShift A := by + letI := E.commShiftFunctor A + have := E.symm.commShift_of_functor A + exact inferInstanceAs (E.symm.symm.CommShift A) + +end Equivalence + +end CategoryTheory From 220ed18e857d7a8d1e9192b899b389c5a88be67b Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 19 Dec 2024 19:41:22 +0000 Subject: [PATCH 791/829] chore(Fintype): move `nonempty_fintype` (#20053) This implication doesn't need the axiom of choice or `Fintype.card`. Currently, it depends on `Classical.choice` because `Fin.fintype` depends on it (for no good reason). --- Mathlib/Algebra/Order/Field/Pi.lean | 2 +- Mathlib/Data/Finite/Defs.lean | 4 ++++ Mathlib/Data/Finset/PiInduction.lean | 2 +- Mathlib/Data/Fintype/Basic.lean | 5 +++++ Mathlib/Data/Fintype/Card.lean | 9 +-------- Mathlib/Data/Fintype/Lattice.lean | 2 +- 6 files changed, 13 insertions(+), 11 deletions(-) diff --git a/Mathlib/Algebra/Order/Field/Pi.lean b/Mathlib/Algebra/Order/Field/Pi.lean index b60f820564ed5..2d128f25cb642 100644 --- a/Mathlib/Algebra/Order/Field/Pi.lean +++ b/Mathlib/Algebra/Order/Field/Pi.lean @@ -5,7 +5,7 @@ Authors: Yaël Dillies -/ import Mathlib.Algebra.Order.Monoid.Defs import Mathlib.Data.Finset.Lattice.Fold -import Mathlib.Data.Fintype.Card +import Mathlib.Data.Fintype.Basic /-! # Lemmas about (finite domain) functions into fields. diff --git a/Mathlib/Data/Finite/Defs.lean b/Mathlib/Data/Finite/Defs.lean index dce5c020d7baf..c1396527a27c7 100644 --- a/Mathlib/Data/Finite/Defs.lean +++ b/Mathlib/Data/Finite/Defs.lean @@ -118,6 +118,10 @@ theorem Function.Bijective.finite_iff {f : α → β} (h : Bijective f) : Finite theorem Finite.ofBijective [Finite α] {f : α → β} (h : Bijective f) : Finite β := h.finite_iff.mp ‹_› +variable (α) in +theorem Finite.nonempty_decidableEq [Finite α] : Nonempty (DecidableEq α) := + let ⟨_n, ⟨e⟩⟩ := Finite.exists_equiv_fin α; ⟨e.decidableEq⟩ + instance [Finite α] : Finite (PLift α) := Finite.of_equiv α Equiv.plift.symm diff --git a/Mathlib/Data/Finset/PiInduction.lean b/Mathlib/Data/Finset/PiInduction.lean index 14801945c227c..36eed999b59bf 100644 --- a/Mathlib/Data/Finset/PiInduction.lean +++ b/Mathlib/Data/Finset/PiInduction.lean @@ -5,7 +5,7 @@ Authors: Yury Kudryashov -/ import Mathlib.Data.Finset.Max import Mathlib.Data.Finset.Sigma -import Mathlib.Data.Fintype.Card +import Mathlib.Data.Fintype.Basic /-! # Induction principles for `∀ i, Finset (α i)` diff --git a/Mathlib/Data/Fintype/Basic.lean b/Mathlib/Data/Fintype/Basic.lean index d381a6d94e4b9..9b443d8cbf4da 100644 --- a/Mathlib/Data/Fintype/Basic.lean +++ b/Mathlib/Data/Fintype/Basic.lean @@ -777,6 +777,11 @@ instance Fin.fintype (n : ℕ) : Fintype (Fin n) := theorem Fin.univ_def (n : ℕ) : (univ : Finset (Fin n)) = ⟨List.finRange n, List.nodup_finRange n⟩ := rfl +/-- See also `nonempty_encodable`, `nonempty_denumerable`. -/ +theorem nonempty_fintype (α : Type*) [Finite α] : Nonempty (Fintype α) := by + rcases Finite.exists_equiv_fin α with ⟨n, ⟨e⟩⟩ + exact ⟨.ofEquiv _ e.symm⟩ + @[simp] theorem List.toFinset_finRange (n : ℕ) : (List.finRange n).toFinset = Finset.univ := by ext; simp diff --git a/Mathlib/Data/Fintype/Card.lean b/Mathlib/Data/Fintype/Card.lean index 0f1f730bb08da..5075621a3a24a 100644 --- a/Mathlib/Data/Fintype/Card.lean +++ b/Mathlib/Data/Fintype/Card.lean @@ -378,14 +378,7 @@ instance (priority := 900) Finite.of_fintype (α : Type*) [Fintype α] : Finite Fintype.finite ‹_› theorem finite_iff_nonempty_fintype (α : Type*) : Finite α ↔ Nonempty (Fintype α) := - ⟨fun h => - let ⟨_k, ⟨e⟩⟩ := @Finite.exists_equiv_fin α h - ⟨Fintype.ofEquiv _ e.symm⟩, - fun ⟨_⟩ => inferInstance⟩ - -/-- See also `nonempty_encodable`, `nonempty_denumerable`. -/ -theorem nonempty_fintype (α : Type*) [Finite α] : Nonempty (Fintype α) := - (finite_iff_nonempty_fintype α).mp ‹_› + ⟨fun _ => nonempty_fintype α, fun ⟨_⟩ => inferInstance⟩ /-- Noncomputably get a `Fintype` instance from a `Finite` instance. This is not an instance because we want `Fintype` instances to be useful for computations. -/ diff --git a/Mathlib/Data/Fintype/Lattice.lean b/Mathlib/Data/Fintype/Lattice.lean index 64865d35bb0c7..e85b48a5901bd 100644 --- a/Mathlib/Data/Fintype/Lattice.lean +++ b/Mathlib/Data/Fintype/Lattice.lean @@ -3,8 +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.Data.Fintype.Card import Mathlib.Data.Finset.Max +import Mathlib.Data.Fintype.Basic /-! # Lemmas relating fintypes and order/lattice structure. From c9f97ab5091a2a3ac9eaba5745ee293ddfee8dde Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Thu, 19 Dec 2024 20:34:31 +0000 Subject: [PATCH 792/829] chore(Topology/Algebra/Module): split large file (#20036) Split up `Mathlib.Topology.Algebra.Module.Basic` (> 2500 lines) into several smaller files: - `Mathlib.Topology.Algebra.Module.Basic`: topological modules and submodules - `Mathlib.Topology.Algebra.Module.LinearMap`: continuous linear maps - `Mathlib.Topology.Algebra.Module.LinearMapPiProd`: pi and prod constructions - `Mathlib.Topology.Algebra.Module.Equiv`: theory of ContinuousLinearEquivs I found it hard to find natural splitting points, so two of the resulting files (LinearMap and Equiv) are still slightly over 1000 lines. --- Mathlib.lean | 3 + Mathlib/Analysis/Convex/Exposed.lean | 2 +- .../InnerProductSpace/LinearPMap.lean | 2 +- Mathlib/Analysis/LocallyConvex/Bounded.lean | 1 - .../Normed/Operator/ContinuousLinearMap.lean | 2 +- .../Normed/Operator/LinearIsometry.lean | 2 +- Mathlib/Analysis/NormedSpace/Real.lean | 1 - .../AffineSpace/ContinuousAffineEquiv.lean | 2 +- Mathlib/MeasureTheory/Group/Measure.lean | 1 + Mathlib/Topology/Algebra/Algebra.lean | 2 +- .../Topology/Algebra/ContinuousAffineMap.lean | 2 +- .../Topology/Algebra/InfiniteSum/Module.lean | 2 +- .../Algebra/Module/Alternating/Basic.lean | 1 + Mathlib/Topology/Algebra/Module/Basic.lean | 2345 +---------------- .../Topology/Algebra/Module/Determinant.lean | 2 +- Mathlib/Topology/Algebra/Module/Equiv.lean | 1035 ++++++++ .../Topology/Algebra/Module/LinearMap.lean | 1080 ++++++++ .../Algebra/Module/LinearMapPiProd.lean | 299 +++ .../Algebra/Module/ModuleTopology.lean | 2 +- .../Algebra/Module/Multilinear/Basic.lean | 2 +- Mathlib/Topology/Algebra/Module/Star.lean | 2 +- .../Algebra/Module/StrongTopology.lean | 1 + Mathlib/Topology/Algebra/Module/WeakDual.lean | 2 +- .../Algebra/Nonarchimedean/Bases.lean | 5 +- .../Algebra/SeparationQuotient/Basic.lean | 2 +- .../Topology/Algebra/UniformFilterBasis.lean | 1 + Mathlib/Topology/ContinuousMap/Algebra.lean | 2 +- .../Topology/Instances/RealVectorSpace.lean | 2 +- Mathlib/Topology/Instances/TrivSqZeroExt.lean | 2 +- 29 files changed, 2468 insertions(+), 2339 deletions(-) create mode 100644 Mathlib/Topology/Algebra/Module/Equiv.lean create mode 100644 Mathlib/Topology/Algebra/Module/LinearMap.lean create mode 100644 Mathlib/Topology/Algebra/Module/LinearMapPiProd.lean diff --git a/Mathlib.lean b/Mathlib.lean index fc9b7328725b2..f04515a65fed1 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -5014,7 +5014,10 @@ import Mathlib.Topology.Algebra.Module.Basic import Mathlib.Topology.Algebra.Module.Cardinality import Mathlib.Topology.Algebra.Module.CharacterSpace import Mathlib.Topology.Algebra.Module.Determinant +import Mathlib.Topology.Algebra.Module.Equiv import Mathlib.Topology.Algebra.Module.FiniteDimension +import Mathlib.Topology.Algebra.Module.LinearMap +import Mathlib.Topology.Algebra.Module.LinearMapPiProd import Mathlib.Topology.Algebra.Module.LinearPMap import Mathlib.Topology.Algebra.Module.LocallyConvex import Mathlib.Topology.Algebra.Module.ModuleTopology diff --git a/Mathlib/Analysis/Convex/Exposed.lean b/Mathlib/Analysis/Convex/Exposed.lean index 064c05aff5bdd..c3d804c5b270b 100644 --- a/Mathlib/Analysis/Convex/Exposed.lean +++ b/Mathlib/Analysis/Convex/Exposed.lean @@ -5,7 +5,7 @@ Authors: Yaël Dillies, Bhavik Mehta -/ import Mathlib.Analysis.Convex.Extreme import Mathlib.Analysis.Convex.Function -import Mathlib.Topology.Algebra.Module.Basic +import Mathlib.Topology.Algebra.Module.LinearMap import Mathlib.Topology.Order.OrderClosed /-! diff --git a/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean b/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean index 8d89fc6d5f0b7..0becfb9ef8a29 100644 --- a/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean +++ b/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Moritz Doll -/ import Mathlib.Analysis.InnerProductSpace.Adjoint -import Mathlib.Topology.Algebra.Module.Basic +import Mathlib.Topology.Algebra.Module.Equiv /-! diff --git a/Mathlib/Analysis/LocallyConvex/Bounded.lean b/Mathlib/Analysis/LocallyConvex/Bounded.lean index 1d3869b828bc5..794cdafedf859 100644 --- a/Mathlib/Analysis/LocallyConvex/Bounded.lean +++ b/Mathlib/Analysis/LocallyConvex/Bounded.lean @@ -11,7 +11,6 @@ import Mathlib.LinearAlgebra.Basis.VectorSpace import Mathlib.Topology.Bornology.Basic import Mathlib.Topology.Algebra.UniformGroup.Basic import Mathlib.Topology.UniformSpace.Cauchy -import Mathlib.Topology.Algebra.Module.Basic /-! # Von Neumann Boundedness diff --git a/Mathlib/Analysis/Normed/Operator/ContinuousLinearMap.lean b/Mathlib/Analysis/Normed/Operator/ContinuousLinearMap.lean index 8b61db66f18eb..3808374cf2a46 100644 --- a/Mathlib/Analysis/Normed/Operator/ContinuousLinearMap.lean +++ b/Mathlib/Analysis/Normed/Operator/ContinuousLinearMap.lean @@ -6,7 +6,7 @@ Authors: Jan-David Salchow, Sébastien Gouëzel, Jean Lo import Mathlib.Analysis.Normed.Group.Uniform import Mathlib.Analysis.Normed.MulAction import Mathlib.LinearAlgebra.DFinsupp -import Mathlib.Topology.Algebra.Module.Basic +import Mathlib.Topology.Algebra.Module.Equiv /-! # Constructions of continuous linear maps between (semi-)normed spaces diff --git a/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean index 7628196bc3d17..a008d3d807bd5 100644 --- a/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean +++ b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean @@ -9,7 +9,7 @@ import Mathlib.Analysis.Normed.Group.Submodule import Mathlib.Analysis.Normed.Group.Uniform import Mathlib.LinearAlgebra.Basis.Defs import Mathlib.LinearAlgebra.DFinsupp -import Mathlib.Topology.Algebra.Module.Basic +import Mathlib.Topology.Algebra.Module.Equiv /-! # (Semi-)linear isometries diff --git a/Mathlib/Analysis/NormedSpace/Real.lean b/Mathlib/Analysis/NormedSpace/Real.lean index aed2117e262d0..d707ae75f4090 100644 --- a/Mathlib/Analysis/NormedSpace/Real.lean +++ b/Mathlib/Analysis/NormedSpace/Real.lean @@ -5,7 +5,6 @@ Authors: Yury Kudryashov, Patrick Massot, Eric Wieser, Yaël Dillies -/ import Mathlib.Analysis.Normed.Module.Basic import Mathlib.LinearAlgebra.Basis.VectorSpace -import Mathlib.Topology.Algebra.Module.Basic /-! # Basic facts about real (semi)normed spaces diff --git a/Mathlib/LinearAlgebra/AffineSpace/ContinuousAffineEquiv.lean b/Mathlib/LinearAlgebra/AffineSpace/ContinuousAffineEquiv.lean index 5631cbfd7acca..d8d068602a7f5 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/ContinuousAffineEquiv.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/ContinuousAffineEquiv.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Rothgang -/ import Mathlib.LinearAlgebra.AffineSpace.AffineEquiv -import Mathlib.Topology.Algebra.Module.Basic +import Mathlib.Topology.Algebra.Module.Equiv /-! # Continuous affine equivalences diff --git a/Mathlib/MeasureTheory/Group/Measure.lean b/Mathlib/MeasureTheory/Group/Measure.lean index 066a497a06598..1c5dbcc1d6ff3 100644 --- a/Mathlib/MeasureTheory/Group/Measure.lean +++ b/Mathlib/MeasureTheory/Group/Measure.lean @@ -5,6 +5,7 @@ Authors: Floris van Doorn -/ import Mathlib.MeasureTheory.Group.Action import Mathlib.MeasureTheory.Measure.Prod +import Mathlib.Topology.Algebra.Module.Equiv import Mathlib.Topology.ContinuousMap.CocompactMap /-! diff --git a/Mathlib/Topology/Algebra/Algebra.lean b/Mathlib/Topology/Algebra/Algebra.lean index 9e7f97fa9ee48..c014fccd8e079 100644 --- a/Mathlib/Topology/Algebra/Algebra.lean +++ b/Mathlib/Topology/Algebra/Algebra.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison, Antoine Chambert-Loir, María Inés de Frutos-Fernández -/ import Mathlib.Algebra.Algebra.Subalgebra.Basic -import Mathlib.Topology.Algebra.Module.Basic import Mathlib.RingTheory.Adjoin.Basic +import Mathlib.Topology.Algebra.Module.LinearMap /-! # Topological (sub)algebras diff --git a/Mathlib/Topology/Algebra/ContinuousAffineMap.lean b/Mathlib/Topology/Algebra/ContinuousAffineMap.lean index 0c94ea673275e..61a9a825016af 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.Algebra.Module.Basic +import Mathlib.Topology.Algebra.Module.LinearMap /-! # Continuous affine maps. diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Module.lean b/Mathlib/Topology/Algebra/InfiniteSum/Module.lean index e8ef8f1cd13dc..d328ba7706cd9 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Module.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Module.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Heather Macbeth, Yury Kudryashov, Frédéric Dupuis -/ import Mathlib.Topology.Algebra.InfiniteSum.Constructions -import Mathlib.Topology.Algebra.Module.Basic +import Mathlib.Topology.Algebra.Module.Equiv /-! # Infinite sums in topological vector spaces -/ diff --git a/Mathlib/Topology/Algebra/Module/Alternating/Basic.lean b/Mathlib/Topology/Algebra/Module/Alternating/Basic.lean index d21554798480b..32ed0fe1db1d8 100644 --- a/Mathlib/Topology/Algebra/Module/Alternating/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Alternating/Basic.lean @@ -5,6 +5,7 @@ Authors: Yury Kudryashov, Heather Macbeth, Sébastien Gouëzel -/ import Mathlib.LinearAlgebra.Alternating.Basic import Mathlib.LinearAlgebra.BilinearMap +import Mathlib.Topology.Algebra.Module.Equiv import Mathlib.Topology.Algebra.Module.Multilinear.Basic /-! diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index 193f14fc095d3..40a95b94007f4 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -5,24 +5,18 @@ Authors: Jan-David Salchow, Sébastien Gouëzel, Jean Lo, Yury Kudryashov, Fréd Heather Macbeth -/ import Mathlib.Algebra.Module.Opposite -import Mathlib.LinearAlgebra.Projection import Mathlib.Topology.Algebra.Group.Quotient import Mathlib.Topology.Algebra.Ring.Basic -import Mathlib.Topology.Algebra.UniformGroup.Defs import Mathlib.Topology.UniformSpace.UniformEmbedding import Mathlib.LinearAlgebra.Finsupp.LinearCombination +import Mathlib.LinearAlgebra.Pi +import Mathlib.Order.OmegaCompletePartialOrder +import Mathlib.LinearAlgebra.Quotient.Defs /-! -# Theory of topological modules and continuous linear maps. +# Theory of topological modules We use the class `ContinuousSMul` for topological (semi) modules and topological vector spaces. - -In this file we define continuous (semi-)linear maps, as semilinear maps between topological -modules which are continuous. The set of continuous semilinear maps between the topological -`R₁`-module `M` and `R₂`-module `M₂` with respect to the `RingHom` `σ` is denoted by `M →SL[σ] M₂`. -Plain linear maps are denoted by `M →L[R] M₂` and star-linear maps by `M →L⋆[R] M₂`. - -The corresponding notation for equivalences is `M ≃SL[σ] M₂`, `M ≃L[R] M₂` and `M ≃L⋆[R] M₂`. -/ assert_not_exists Star.star @@ -216,105 +210,6 @@ theorem LinearMap.continuous_on_pi {ι : Type*} {R : Type*} {M : Type*} [Finite end Pi -/-- Continuous linear maps between modules. We only put the type classes that are necessary for the -definition, although in applications `M` and `M₂` will be topological modules over the topological -ring `R`. -/ -structure ContinuousLinearMap {R : Type*} {S : Type*} [Semiring R] [Semiring S] (σ : R →+* S) - (M : Type*) [TopologicalSpace M] [AddCommMonoid M] (M₂ : Type*) [TopologicalSpace M₂] - [AddCommMonoid M₂] [Module R M] [Module S M₂] extends M →ₛₗ[σ] M₂ where - cont : Continuous toFun := by continuity - -attribute [inherit_doc ContinuousLinearMap] ContinuousLinearMap.cont - -@[inherit_doc] -notation:25 M " →SL[" σ "] " M₂ => ContinuousLinearMap σ M M₂ - -@[inherit_doc] -notation:25 M " →L[" R "] " M₂ => ContinuousLinearMap (RingHom.id R) M M₂ - -/-- `ContinuousSemilinearMapClass F σ M M₂` asserts `F` is a type of bundled continuous -`σ`-semilinear maps `M → M₂`. See also `ContinuousLinearMapClass F R M M₂` for the case where -`σ` is the identity map on `R`. A map `f` between an `R`-module and an `S`-module over a ring -homomorphism `σ : R →+* S` is semilinear if it satisfies the two properties `f (x + y) = f x + f y` -and `f (c • x) = (σ c) • f x`. -/ -class ContinuousSemilinearMapClass (F : Type*) {R S : outParam Type*} [Semiring R] [Semiring S] - (σ : outParam <| R →+* S) (M : outParam Type*) [TopologicalSpace M] [AddCommMonoid M] - (M₂ : outParam Type*) [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R M] - [Module S M₂] [FunLike F M M₂] - extends SemilinearMapClass F σ M M₂, ContinuousMapClass F M M₂ : Prop - -/-- `ContinuousLinearMapClass F R M M₂` asserts `F` is a type of bundled continuous -`R`-linear maps `M → M₂`. This is an abbreviation for -`ContinuousSemilinearMapClass F (RingHom.id R) M M₂`. -/ -abbrev ContinuousLinearMapClass (F : Type*) (R : outParam Type*) [Semiring R] - (M : outParam Type*) [TopologicalSpace M] [AddCommMonoid M] (M₂ : outParam Type*) - [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R M] [Module R M₂] [FunLike F M M₂] := - ContinuousSemilinearMapClass F (RingHom.id R) M M₂ - -/-- Continuous linear equivalences between modules. We only put the type classes that are necessary -for the definition, although in applications `M` and `M₂` will be topological modules over the -topological semiring `R`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not ported yet; was @[nolint has_nonempty_instance] -structure ContinuousLinearEquiv {R : Type*} {S : Type*} [Semiring R] [Semiring S] (σ : R →+* S) - {σ' : S →+* R} [RingHomInvPair σ σ'] [RingHomInvPair σ' σ] (M : Type*) [TopologicalSpace M] - [AddCommMonoid M] (M₂ : Type*) [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R M] - [Module S M₂] extends M ≃ₛₗ[σ] M₂ where - continuous_toFun : Continuous toFun := by continuity - continuous_invFun : Continuous invFun := by continuity - -attribute [inherit_doc ContinuousLinearEquiv] ContinuousLinearEquiv.continuous_toFun -ContinuousLinearEquiv.continuous_invFun - -@[inherit_doc] -notation:50 M " ≃SL[" σ "] " M₂ => ContinuousLinearEquiv σ M M₂ - -@[inherit_doc] -notation:50 M " ≃L[" R "] " M₂ => ContinuousLinearEquiv (RingHom.id R) M M₂ - -/-- `ContinuousSemilinearEquivClass F σ M M₂` asserts `F` is a type of bundled continuous -`σ`-semilinear equivs `M → M₂`. See also `ContinuousLinearEquivClass F R M M₂` for the case -where `σ` is the identity map on `R`. A map `f` between an `R`-module and an `S`-module over a ring -homomorphism `σ : R →+* S` is semilinear if it satisfies the two properties `f (x + y) = f x + f y` -and `f (c • x) = (σ c) • f x`. -/ -class ContinuousSemilinearEquivClass (F : Type*) {R : outParam Type*} {S : outParam Type*} - [Semiring R] [Semiring S] (σ : outParam <| R →+* S) {σ' : outParam <| S →+* R} - [RingHomInvPair σ σ'] [RingHomInvPair σ' σ] (M : outParam Type*) [TopologicalSpace M] - [AddCommMonoid M] (M₂ : outParam Type*) [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R M] - [Module S M₂] [EquivLike F M M₂] extends SemilinearEquivClass F σ M M₂ : Prop where - map_continuous : ∀ f : F, Continuous f := by continuity - inv_continuous : ∀ f : F, Continuous (EquivLike.inv f) := by continuity - -attribute [inherit_doc ContinuousSemilinearEquivClass] -ContinuousSemilinearEquivClass.map_continuous -ContinuousSemilinearEquivClass.inv_continuous - -/-- `ContinuousLinearEquivClass F σ M M₂` asserts `F` is a type of bundled continuous -`R`-linear equivs `M → M₂`. This is an abbreviation for -`ContinuousSemilinearEquivClass F (RingHom.id R) M M₂`. -/ -abbrev ContinuousLinearEquivClass (F : Type*) (R : outParam Type*) [Semiring R] - (M : outParam Type*) [TopologicalSpace M] [AddCommMonoid M] (M₂ : outParam Type*) - [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R M] [Module R M₂] [EquivLike F M M₂] := - ContinuousSemilinearEquivClass F (RingHom.id R) M M₂ - -namespace ContinuousSemilinearEquivClass - -variable (F : Type*) {R : Type*} {S : Type*} [Semiring R] [Semiring S] (σ : R →+* S) - {σ' : S →+* R} [RingHomInvPair σ σ'] [RingHomInvPair σ' σ] - (M : Type*) [TopologicalSpace M] [AddCommMonoid M] - (M₂ : Type*) [TopologicalSpace M₂] [AddCommMonoid M₂] - [Module R M] [Module S M₂] - --- `σ'` becomes a metavariable, but it's OK since it's an outparam -instance (priority := 100) continuousSemilinearMapClass [EquivLike F M M₂] - [s : ContinuousSemilinearEquivClass F σ M M₂] : ContinuousSemilinearMapClass F σ M M₂ := - { s with } - -instance (priority := 100) [EquivLike F M M₂] - [s : ContinuousSemilinearEquivClass F σ M M₂] : HomeomorphClass F M M₂ := - { s with } - -end ContinuousSemilinearEquivClass - section PointwiseLimits variable {M₁ M₂ α R S : Type*} [TopologicalSpace M₂] [T2Space M₂] [Semiring R] [Semiring S] @@ -345,2223 +240,37 @@ theorem LinearMap.isClosed_range_coe : IsClosed (Set.range ((↑) : (M₁ →ₛ end PointwiseLimits -namespace ContinuousLinearMap - -section Semiring - -/-! -### Properties that hold for non-necessarily commutative semirings. --/ - -variable {R₁ : Type*} {R₂ : Type*} {R₃ : Type*} [Semiring R₁] [Semiring R₂] [Semiring R₃] - {σ₁₂ : R₁ →+* R₂} {σ₂₃ : R₂ →+* R₃} {σ₁₃ : R₁ →+* R₃} {M₁ : Type*} [TopologicalSpace M₁] - [AddCommMonoid M₁] {M'₁ : Type*} [TopologicalSpace M'₁] [AddCommMonoid M'₁] {M₂ : Type*} - [TopologicalSpace M₂] [AddCommMonoid M₂] {M₃ : Type*} [TopologicalSpace M₃] [AddCommMonoid M₃] - {M₄ : Type*} [TopologicalSpace M₄] [AddCommMonoid M₄] [Module R₁ M₁] [Module R₁ M'₁] - [Module R₂ M₂] [Module R₃ M₃] - -attribute [coe] ContinuousLinearMap.toLinearMap -/-- Coerce continuous linear maps to linear maps. -/ -instance LinearMap.coe : Coe (M₁ →SL[σ₁₂] M₂) (M₁ →ₛₗ[σ₁₂] M₂) := ⟨toLinearMap⟩ - -theorem coe_injective : Function.Injective ((↑) : (M₁ →SL[σ₁₂] M₂) → M₁ →ₛₗ[σ₁₂] M₂) := by - intro f g H - cases f - cases g - congr - -instance funLike : FunLike (M₁ →SL[σ₁₂] M₂) M₁ M₂ where - coe f := f.toLinearMap - coe_injective' _ _ h := coe_injective (DFunLike.coe_injective h) - -instance continuousSemilinearMapClass : - ContinuousSemilinearMapClass (M₁ →SL[σ₁₂] M₂) σ₁₂ M₁ M₂ where - map_add f := map_add f.toLinearMap - map_continuous f := f.2 - map_smulₛₗ f := f.toLinearMap.map_smul' - -theorem coe_mk (f : M₁ →ₛₗ[σ₁₂] M₂) (h) : (mk f h : M₁ →ₛₗ[σ₁₂] M₂) = f := - rfl - -@[simp] -theorem coe_mk' (f : M₁ →ₛₗ[σ₁₂] M₂) (h) : (mk f h : M₁ → M₂) = f := - rfl - -@[continuity, fun_prop] -protected theorem continuous (f : M₁ →SL[σ₁₂] M₂) : Continuous f := - f.2 - -protected theorem uniformContinuous {E₁ E₂ : Type*} [UniformSpace E₁] [UniformSpace E₂] - [AddCommGroup E₁] [AddCommGroup E₂] [Module R₁ E₁] [Module R₂ E₂] [UniformAddGroup E₁] - [UniformAddGroup E₂] (f : E₁ →SL[σ₁₂] E₂) : UniformContinuous f := - uniformContinuous_addMonoidHom_of_continuous f.continuous - -@[simp, norm_cast] -theorem coe_inj {f g : M₁ →SL[σ₁₂] M₂} : (f : M₁ →ₛₗ[σ₁₂] M₂) = g ↔ f = g := - coe_injective.eq_iff - -theorem coeFn_injective : @Function.Injective (M₁ →SL[σ₁₂] M₂) (M₁ → M₂) (↑) := - DFunLike.coe_injective - -/-- See Note [custom simps projection]. We need to specify this projection explicitly in this case, - because it is a composition of multiple projections. -/ -def Simps.apply (h : M₁ →SL[σ₁₂] M₂) : M₁ → M₂ := - h - -/-- See Note [custom simps projection]. -/ -def Simps.coe (h : M₁ →SL[σ₁₂] M₂) : M₁ →ₛₗ[σ₁₂] M₂ := - h - -initialize_simps_projections ContinuousLinearMap (toFun → apply, toLinearMap → coe) - -@[ext] -theorem ext {f g : M₁ →SL[σ₁₂] M₂} (h : ∀ x, f x = g x) : f = g := - DFunLike.ext f g h - -/-- Copy of a `ContinuousLinearMap` with a new `toFun` equal to the old one. Useful to fix -definitional equalities. -/ -protected def copy (f : M₁ →SL[σ₁₂] M₂) (f' : M₁ → M₂) (h : f' = ⇑f) : M₁ →SL[σ₁₂] M₂ where - toLinearMap := f.toLinearMap.copy f' h - cont := show Continuous f' from h.symm ▸ f.continuous - -@[simp] -theorem coe_copy (f : M₁ →SL[σ₁₂] M₂) (f' : M₁ → M₂) (h : f' = ⇑f) : ⇑(f.copy f' h) = f' := - rfl - -theorem copy_eq (f : M₁ →SL[σ₁₂] M₂) (f' : M₁ → M₂) (h : f' = ⇑f) : f.copy f' h = f := - DFunLike.ext' h - -theorem range_coeFn_eq : - Set.range ((⇑) : (M₁ →SL[σ₁₂] M₂) → (M₁ → M₂)) = - {f | Continuous f} ∩ Set.range ((⇑) : (M₁ →ₛₗ[σ₁₂] M₂) → (M₁ → M₂)) := by - ext f - constructor - · rintro ⟨f, rfl⟩ - exact ⟨f.continuous, f, rfl⟩ - · rintro ⟨hfc, f, rfl⟩ - exact ⟨⟨f, hfc⟩, rfl⟩ - --- make some straightforward lemmas available to `simp`. -protected theorem map_zero (f : M₁ →SL[σ₁₂] M₂) : f (0 : M₁) = 0 := - map_zero f - -protected theorem map_add (f : M₁ →SL[σ₁₂] M₂) (x y : M₁) : f (x + y) = f x + f y := - map_add f x y - -@[simp] -protected theorem map_smulₛₗ (f : M₁ →SL[σ₁₂] M₂) (c : R₁) (x : M₁) : f (c • x) = σ₁₂ c • f x := - (toLinearMap _).map_smulₛₗ _ _ - -protected theorem map_smul [Module R₁ M₂] (f : M₁ →L[R₁] M₂) (c : R₁) (x : M₁) : - f (c • x) = c • f x := by simp only [RingHom.id_apply, ContinuousLinearMap.map_smulₛₗ] - -@[simp] -theorem map_smul_of_tower {R S : Type*} [Semiring S] [SMul R M₁] [Module S M₁] [SMul R M₂] - [Module S M₂] [LinearMap.CompatibleSMul M₁ M₂ R S] (f : M₁ →L[S] M₂) (c : R) (x : M₁) : - f (c • x) = c • f x := - LinearMap.CompatibleSMul.map_smul (f : M₁ →ₗ[S] M₂) c x - -@[simp, norm_cast] -theorem coe_coe (f : M₁ →SL[σ₁₂] M₂) : ⇑(f : M₁ →ₛₗ[σ₁₂] M₂) = f := - rfl - -@[ext] -theorem ext_ring [TopologicalSpace R₁] {f g : R₁ →L[R₁] M₁} (h : f 1 = g 1) : f = g := - coe_inj.1 <| LinearMap.ext_ring h - -/-- If two continuous linear maps are equal on a set `s`, then they are equal on the closure -of the `Submodule.span` of this set. -/ -theorem eqOn_closure_span [T2Space M₂] {s : Set M₁} {f g : M₁ →SL[σ₁₂] M₂} (h : Set.EqOn f g s) : - Set.EqOn f g (closure (Submodule.span R₁ s : Set M₁)) := - (LinearMap.eqOn_span' h).closure f.continuous g.continuous - -/-- If the submodule generated by a set `s` is dense in the ambient module, then two continuous -linear maps equal on `s` are equal. -/ -theorem ext_on [T2Space M₂] {s : Set M₁} (hs : Dense (Submodule.span R₁ s : Set M₁)) - {f g : M₁ →SL[σ₁₂] M₂} (h : Set.EqOn f g s) : f = g := - ext fun x => eqOn_closure_span h (hs x) - -/-- Under a continuous linear map, the image of the `TopologicalClosure` of a submodule is -contained in the `TopologicalClosure` of its image. -/ -theorem _root_.Submodule.topologicalClosure_map [RingHomSurjective σ₁₂] [TopologicalSpace R₁] - [TopologicalSpace R₂] [ContinuousSMul R₁ M₁] [ContinuousAdd M₁] [ContinuousSMul R₂ M₂] - [ContinuousAdd M₂] (f : M₁ →SL[σ₁₂] M₂) (s : Submodule R₁ M₁) : - s.topologicalClosure.map (f : M₁ →ₛₗ[σ₁₂] M₂) ≤ - (s.map (f : M₁ →ₛₗ[σ₁₂] M₂)).topologicalClosure := - image_closure_subset_closure_image f.continuous - -/-- Under a dense continuous linear map, a submodule whose `TopologicalClosure` is `⊤` is sent to -another such submodule. That is, the image of a dense set under a map with dense range is dense. --/ -theorem _root_.DenseRange.topologicalClosure_map_submodule [RingHomSurjective σ₁₂] - [TopologicalSpace R₁] [TopologicalSpace R₂] [ContinuousSMul R₁ M₁] [ContinuousAdd M₁] - [ContinuousSMul R₂ M₂] [ContinuousAdd M₂] {f : M₁ →SL[σ₁₂] M₂} (hf' : DenseRange f) - {s : Submodule R₁ M₁} (hs : s.topologicalClosure = ⊤) : - (s.map (f : M₁ →ₛₗ[σ₁₂] M₂)).topologicalClosure = ⊤ := by - rw [SetLike.ext'_iff] at hs ⊢ - simp only [Submodule.topologicalClosure_coe, Submodule.top_coe, ← dense_iff_closure_eq] at hs ⊢ - exact hf'.dense_image f.continuous hs - -section SMulMonoid - -variable {S₂ T₂ : Type*} [Monoid S₂] [Monoid T₂] -variable [DistribMulAction S₂ M₂] [SMulCommClass R₂ S₂ M₂] [ContinuousConstSMul S₂ M₂] -variable [DistribMulAction T₂ M₂] [SMulCommClass R₂ T₂ M₂] [ContinuousConstSMul T₂ M₂] - -instance instSMul : SMul S₂ (M₁ →SL[σ₁₂] M₂) where - smul c f := ⟨c • (f : M₁ →ₛₗ[σ₁₂] M₂), (f.2.const_smul _ : Continuous fun x => c • f x)⟩ - -instance mulAction : MulAction S₂ (M₁ →SL[σ₁₂] M₂) where - one_smul _f := ext fun _x => one_smul _ _ - mul_smul _a _b _f := ext fun _x => mul_smul _ _ _ - -theorem smul_apply (c : S₂) (f : M₁ →SL[σ₁₂] M₂) (x : M₁) : (c • f) x = c • f x := - rfl - -@[simp, norm_cast] -theorem coe_smul (c : S₂) (f : M₁ →SL[σ₁₂] M₂) : - ↑(c • f) = c • (f : M₁ →ₛₗ[σ₁₂] M₂) := - rfl - -@[simp, norm_cast] -theorem coe_smul' (c : S₂) (f : M₁ →SL[σ₁₂] M₂) : - ↑(c • f) = c • (f : M₁ → M₂) := - rfl - -instance isScalarTower [SMul S₂ T₂] [IsScalarTower S₂ T₂ M₂] : - IsScalarTower S₂ T₂ (M₁ →SL[σ₁₂] M₂) := - ⟨fun a b f => ext fun x => smul_assoc a b (f x)⟩ - -instance smulCommClass [SMulCommClass S₂ T₂ M₂] : SMulCommClass S₂ T₂ (M₁ →SL[σ₁₂] M₂) := - ⟨fun a b f => ext fun x => smul_comm a b (f x)⟩ - -end SMulMonoid - -/-- The continuous map that is constantly zero. -/ -instance zero : Zero (M₁ →SL[σ₁₂] M₂) := - ⟨⟨0, continuous_zero⟩⟩ - -instance inhabited : Inhabited (M₁ →SL[σ₁₂] M₂) := - ⟨0⟩ - -@[simp] -theorem default_def : (default : M₁ →SL[σ₁₂] M₂) = 0 := - rfl - -@[simp] -theorem zero_apply (x : M₁) : (0 : M₁ →SL[σ₁₂] M₂) x = 0 := - rfl - -@[simp, norm_cast] -theorem coe_zero : ((0 : M₁ →SL[σ₁₂] M₂) : M₁ →ₛₗ[σ₁₂] M₂) = 0 := - rfl - -/- no simp attribute on the next line as simp does not always simplify `0 x` to `0` -when `0` is the zero function, while it does for the zero continuous linear map, -and this is the most important property we care about. -/ -@[norm_cast] -theorem coe_zero' : ⇑(0 : M₁ →SL[σ₁₂] M₂) = 0 := - rfl - -instance uniqueOfLeft [Subsingleton M₁] : Unique (M₁ →SL[σ₁₂] M₂) := - coe_injective.unique - -instance uniqueOfRight [Subsingleton M₂] : Unique (M₁ →SL[σ₁₂] M₂) := - coe_injective.unique - -theorem exists_ne_zero {f : M₁ →SL[σ₁₂] M₂} (hf : f ≠ 0) : ∃ x, f x ≠ 0 := by - by_contra! h - exact hf (ContinuousLinearMap.ext h) - -section - -variable (R₁ M₁) - -/-- the identity map as a continuous linear map. -/ -def id : M₁ →L[R₁] M₁ := - ⟨LinearMap.id, continuous_id⟩ - -end - -instance one : One (M₁ →L[R₁] M₁) := - ⟨id R₁ M₁⟩ - -theorem one_def : (1 : M₁ →L[R₁] M₁) = id R₁ M₁ := - rfl - -theorem id_apply (x : M₁) : id R₁ M₁ x = x := - rfl - -@[simp, norm_cast] -theorem coe_id : (id R₁ M₁ : M₁ →ₗ[R₁] M₁) = LinearMap.id := - rfl - -@[simp, norm_cast] -theorem coe_id' : ⇑(id R₁ M₁) = _root_.id := - rfl - -@[simp, norm_cast] -theorem coe_eq_id {f : M₁ →L[R₁] M₁} : (f : M₁ →ₗ[R₁] M₁) = LinearMap.id ↔ f = id _ _ := by - rw [← coe_id, coe_inj] - -@[simp] -theorem one_apply (x : M₁) : (1 : M₁ →L[R₁] M₁) x = x := - rfl - -instance [Nontrivial M₁] : Nontrivial (M₁ →L[R₁] M₁) := - ⟨0, 1, fun e ↦ - have ⟨x, hx⟩ := exists_ne (0 : M₁); hx (by simpa using DFunLike.congr_fun e.symm x)⟩ - -section Add - -variable [ContinuousAdd M₂] - -instance add : Add (M₁ →SL[σ₁₂] M₂) := - ⟨fun f g => ⟨f + g, f.2.add g.2⟩⟩ - -@[simp] -theorem add_apply (f g : M₁ →SL[σ₁₂] M₂) (x : M₁) : (f + g) x = f x + g x := - rfl - -@[simp, norm_cast] -theorem coe_add (f g : M₁ →SL[σ₁₂] M₂) : (↑(f + g) : M₁ →ₛₗ[σ₁₂] M₂) = f + g := - rfl - -@[norm_cast] -theorem coe_add' (f g : M₁ →SL[σ₁₂] M₂) : ⇑(f + g) = f + g := - rfl - -instance addCommMonoid : AddCommMonoid (M₁ →SL[σ₁₂] M₂) where - zero_add := by - intros - ext - apply_rules [zero_add, add_assoc, add_zero, neg_add_cancel, add_comm] - add_zero := by - intros - ext - apply_rules [zero_add, add_assoc, add_zero, neg_add_cancel, add_comm] - add_comm := by - intros - ext - apply_rules [zero_add, add_assoc, add_zero, neg_add_cancel, add_comm] - add_assoc := by - intros - ext - apply_rules [zero_add, add_assoc, add_zero, neg_add_cancel, add_comm] - nsmul := (· • ·) - nsmul_zero f := by - ext - simp - nsmul_succ n f := by - ext - simp [add_smul] - -@[simp, norm_cast] -theorem coe_sum {ι : Type*} (t : Finset ι) (f : ι → M₁ →SL[σ₁₂] M₂) : - ↑(∑ d ∈ t, f d) = (∑ d ∈ t, f d : M₁ →ₛₗ[σ₁₂] M₂) := - map_sum (AddMonoidHom.mk ⟨((↑) : (M₁ →SL[σ₁₂] M₂) → M₁ →ₛₗ[σ₁₂] M₂), rfl⟩ fun _ _ => rfl) _ _ - -@[simp, norm_cast] -theorem coe_sum' {ι : Type*} (t : Finset ι) (f : ι → M₁ →SL[σ₁₂] M₂) : - ⇑(∑ d ∈ t, f d) = ∑ d ∈ t, ⇑(f d) := by simp only [← coe_coe, coe_sum, LinearMap.coeFn_sum] - -theorem sum_apply {ι : Type*} (t : Finset ι) (f : ι → M₁ →SL[σ₁₂] M₂) (b : M₁) : - (∑ d ∈ t, f d) b = ∑ d ∈ t, f d b := by simp only [coe_sum', Finset.sum_apply] - -end Add - -variable [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] - -/-- Composition of bounded linear maps. -/ -def comp (g : M₂ →SL[σ₂₃] M₃) (f : M₁ →SL[σ₁₂] M₂) : M₁ →SL[σ₁₃] M₃ := - ⟨(g : M₂ →ₛₗ[σ₂₃] M₃).comp (f : M₁ →ₛₗ[σ₁₂] M₂), g.2.comp f.2⟩ - -@[inherit_doc comp] -infixr:80 " ∘L " => - @ContinuousLinearMap.comp _ _ _ _ _ _ (RingHom.id _) (RingHom.id _) (RingHom.id _) _ _ _ _ _ _ _ _ - _ _ _ _ RingHomCompTriple.ids - -@[simp, norm_cast] -theorem coe_comp (h : M₂ →SL[σ₂₃] M₃) (f : M₁ →SL[σ₁₂] M₂) : - (h.comp f : M₁ →ₛₗ[σ₁₃] M₃) = (h : M₂ →ₛₗ[σ₂₃] M₃).comp (f : M₁ →ₛₗ[σ₁₂] M₂) := - rfl - -@[simp, norm_cast] -theorem coe_comp' (h : M₂ →SL[σ₂₃] M₃) (f : M₁ →SL[σ₁₂] M₂) : ⇑(h.comp f) = h ∘ f := - rfl - -theorem comp_apply (g : M₂ →SL[σ₂₃] M₃) (f : M₁ →SL[σ₁₂] M₂) (x : M₁) : (g.comp f) x = g (f x) := - rfl - -@[simp] -theorem comp_id (f : M₁ →SL[σ₁₂] M₂) : f.comp (id R₁ M₁) = f := - ext fun _x => rfl - -@[simp] -theorem id_comp (f : M₁ →SL[σ₁₂] M₂) : (id R₂ M₂).comp f = f := - ext fun _x => rfl - -@[simp] -theorem comp_zero (g : M₂ →SL[σ₂₃] M₃) : g.comp (0 : M₁ →SL[σ₁₂] M₂) = 0 := by - ext - simp - -@[simp] -theorem zero_comp (f : M₁ →SL[σ₁₂] M₂) : (0 : M₂ →SL[σ₂₃] M₃).comp f = 0 := by - ext - simp - -@[simp] -theorem comp_add [ContinuousAdd M₂] [ContinuousAdd M₃] (g : M₂ →SL[σ₂₃] M₃) - (f₁ f₂ : M₁ →SL[σ₁₂] M₂) : g.comp (f₁ + f₂) = g.comp f₁ + g.comp f₂ := by - ext - simp - -@[simp] -theorem add_comp [ContinuousAdd M₃] (g₁ g₂ : M₂ →SL[σ₂₃] M₃) (f : M₁ →SL[σ₁₂] M₂) : - (g₁ + g₂).comp f = g₁.comp f + g₂.comp f := by - ext - simp - -theorem comp_finset_sum {ι : Type*} {s : Finset ι} - [ContinuousAdd M₂] [ContinuousAdd M₃] (g : M₂ →SL[σ₂₃] M₃) - (f : ι → M₁ →SL[σ₁₂] M₂) : g.comp (∑ i ∈ s, f i) = ∑ i ∈ s, g.comp (f i) := by - ext - simp - -theorem finset_sum_comp {ι : Type*} {s : Finset ι} - [ContinuousAdd M₃] (g : ι → M₂ →SL[σ₂₃] M₃) - (f : M₁ →SL[σ₁₂] M₂) : (∑ i ∈ s, g i).comp f = ∑ i ∈ s, (g i).comp f := by - ext - simp only [coe_comp', coe_sum', Function.comp_apply, Finset.sum_apply] - -theorem comp_assoc {R₄ : Type*} [Semiring R₄] [Module R₄ M₄] {σ₁₄ : R₁ →+* R₄} {σ₂₄ : R₂ →+* R₄} - {σ₃₄ : R₃ →+* R₄} [RingHomCompTriple σ₁₃ σ₃₄ σ₁₄] [RingHomCompTriple σ₂₃ σ₃₄ σ₂₄] - [RingHomCompTriple σ₁₂ σ₂₄ σ₁₄] (h : M₃ →SL[σ₃₄] M₄) (g : M₂ →SL[σ₂₃] M₃) (f : M₁ →SL[σ₁₂] M₂) : - (h.comp g).comp f = h.comp (g.comp f) := - rfl - -instance instMul : Mul (M₁ →L[R₁] M₁) := - ⟨comp⟩ - -theorem mul_def (f g : M₁ →L[R₁] M₁) : f * g = f.comp g := - rfl - -@[simp] -theorem coe_mul (f g : M₁ →L[R₁] M₁) : ⇑(f * g) = f ∘ g := - rfl - -theorem mul_apply (f g : M₁ →L[R₁] M₁) (x : M₁) : (f * g) x = f (g x) := - rfl - -instance monoidWithZero : MonoidWithZero (M₁ →L[R₁] M₁) where - mul_zero f := ext fun _ => map_zero f - zero_mul _ := ext fun _ => rfl - mul_one _ := ext fun _ => rfl - one_mul _ := ext fun _ => rfl - mul_assoc _ _ _ := ext fun _ => rfl - -theorem coe_pow (f : M₁ →L[R₁] M₁) (n : ℕ) : ⇑(f ^ n) = f^[n] := - hom_coe_pow _ rfl (fun _ _ ↦ rfl) _ _ - -instance instNatCast [ContinuousAdd M₁] : NatCast (M₁ →L[R₁] M₁) where - natCast n := n • (1 : M₁ →L[R₁] M₁) - -instance semiring [ContinuousAdd M₁] : Semiring (M₁ →L[R₁] M₁) where - __ := ContinuousLinearMap.monoidWithZero - __ := ContinuousLinearMap.addCommMonoid - left_distrib f g h := ext fun x => map_add f (g x) (h x) - right_distrib _ _ _ := ext fun _ => LinearMap.add_apply _ _ _ - toNatCast := instNatCast - natCast_zero := zero_smul ℕ (1 : M₁ →L[R₁] M₁) - natCast_succ n := AddMonoid.nsmul_succ n (1 : M₁ →L[R₁] M₁) - -/-- `ContinuousLinearMap.toLinearMap` as a `RingHom`. -/ -@[simps] -def toLinearMapRingHom [ContinuousAdd M₁] : (M₁ →L[R₁] M₁) →+* M₁ →ₗ[R₁] M₁ where - toFun := toLinearMap - map_zero' := rfl - map_one' := rfl - map_add' _ _ := rfl - map_mul' _ _ := rfl - -@[simp] -theorem natCast_apply [ContinuousAdd M₁] (n : ℕ) (m : M₁) : (↑n : M₁ →L[R₁] M₁) m = n • m := - rfl - -@[simp] -theorem ofNat_apply [ContinuousAdd M₁] (n : ℕ) [n.AtLeastTwo] (m : M₁) : - ((no_index (OfNat.ofNat n) : M₁ →L[R₁] M₁)) m = OfNat.ofNat n • m := - rfl - -section ApplyAction - -variable [ContinuousAdd M₁] - -/-- The tautological action by `M₁ →L[R₁] M₁` on `M`. - -This generalizes `Function.End.applyMulAction`. -/ -instance applyModule : Module (M₁ →L[R₁] M₁) M₁ := - Module.compHom _ toLinearMapRingHom - -@[simp] -protected theorem smul_def (f : M₁ →L[R₁] M₁) (a : M₁) : f • a = f a := - rfl - -/-- `ContinuousLinearMap.applyModule` is faithful. -/ -instance applyFaithfulSMul : FaithfulSMul (M₁ →L[R₁] M₁) M₁ := - ⟨fun {_ _} => ContinuousLinearMap.ext⟩ - -instance applySMulCommClass : SMulCommClass R₁ (M₁ →L[R₁] M₁) M₁ where - smul_comm r e m := (e.map_smul r m).symm - -instance applySMulCommClass' : SMulCommClass (M₁ →L[R₁] M₁) R₁ M₁ where - smul_comm := ContinuousLinearMap.map_smul - -instance continuousConstSMul_apply : ContinuousConstSMul (M₁ →L[R₁] M₁) M₁ := - ⟨ContinuousLinearMap.continuous⟩ - -end ApplyAction - -/-- The cartesian product of two bounded linear maps, as a bounded linear map. -/ -protected def prod [Module R₁ M₂] [Module R₁ M₃] (f₁ : M₁ →L[R₁] M₂) (f₂ : M₁ →L[R₁] M₃) : - M₁ →L[R₁] M₂ × M₃ := - ⟨(f₁ : M₁ →ₗ[R₁] M₂).prod f₂, f₁.2.prod_mk f₂.2⟩ - -@[simp, norm_cast] -theorem coe_prod [Module R₁ M₂] [Module R₁ M₃] (f₁ : M₁ →L[R₁] M₂) (f₂ : M₁ →L[R₁] M₃) : - (f₁.prod f₂ : M₁ →ₗ[R₁] M₂ × M₃) = LinearMap.prod f₁ f₂ := - rfl - -@[simp, norm_cast] -theorem prod_apply [Module R₁ M₂] [Module R₁ M₃] (f₁ : M₁ →L[R₁] M₂) (f₂ : M₁ →L[R₁] M₃) (x : M₁) : - f₁.prod f₂ x = (f₁ x, f₂ x) := - rfl - -section - -variable (R₁ M₁ M₂) - -/-- The left injection into a product is a continuous linear map. -/ -def inl [Module R₁ M₂] : M₁ →L[R₁] M₁ × M₂ := - (id R₁ M₁).prod 0 - -/-- The right injection into a product is a continuous linear map. -/ -def inr [Module R₁ M₂] : M₂ →L[R₁] M₁ × M₂ := - (0 : M₂ →L[R₁] M₁).prod (id R₁ M₂) - -end - -variable {F : Type*} - -@[simp] -theorem inl_apply [Module R₁ M₂] (x : M₁) : inl R₁ M₁ M₂ x = (x, 0) := - rfl - -@[simp] -theorem inr_apply [Module R₁ M₂] (x : M₂) : inr R₁ M₁ M₂ x = (0, x) := - rfl - -@[simp, norm_cast] -theorem coe_inl [Module R₁ M₂] : (inl R₁ M₁ M₂ : M₁ →ₗ[R₁] M₁ × M₂) = LinearMap.inl R₁ M₁ M₂ := - rfl - -@[simp, norm_cast] -theorem coe_inr [Module R₁ M₂] : (inr R₁ M₁ M₂ : M₂ →ₗ[R₁] M₁ × M₂) = LinearMap.inr R₁ M₁ M₂ := - rfl - -theorem isClosed_ker [T1Space M₂] [FunLike F M₁ M₂] [ContinuousSemilinearMapClass F σ₁₂ M₁ M₂] - (f : F) : - IsClosed (ker f : Set M₁) := - continuous_iff_isClosed.1 (map_continuous f) _ isClosed_singleton - -theorem isComplete_ker {M' : Type*} [UniformSpace M'] [CompleteSpace M'] [AddCommMonoid M'] - [Module R₁ M'] [T1Space M₂] [FunLike F M' M₂] [ContinuousSemilinearMapClass F σ₁₂ M' M₂] - (f : F) : - IsComplete (ker f : Set M') := - (isClosed_ker f).isComplete - -instance completeSpace_ker {M' : Type*} [UniformSpace M'] [CompleteSpace M'] - [AddCommMonoid M'] [Module R₁ M'] [T1Space M₂] - [FunLike F M' M₂] [ContinuousSemilinearMapClass F σ₁₂ M' M₂] - (f : F) : CompleteSpace (ker f) := - (isComplete_ker f).completeSpace_coe - -instance completeSpace_eqLocus {M' : Type*} [UniformSpace M'] [CompleteSpace M'] - [AddCommMonoid M'] [Module R₁ M'] [T2Space M₂] - [FunLike F M' M₂] [ContinuousSemilinearMapClass F σ₁₂ M' M₂] - (f g : F) : CompleteSpace (LinearMap.eqLocus f g) := - IsClosed.completeSpace_coe <| isClosed_eq (map_continuous f) (map_continuous g) - -@[simp] -theorem ker_prod [Module R₁ M₂] [Module R₁ M₃] (f : M₁ →L[R₁] M₂) (g : M₁ →L[R₁] M₃) : - ker (f.prod g) = ker f ⊓ ker g := - LinearMap.ker_prod (f : M₁ →ₗ[R₁] M₂) (g : M₁ →ₗ[R₁] M₃) - -/-- Restrict codomain of a continuous linear map. -/ -def codRestrict (f : M₁ →SL[σ₁₂] M₂) (p : Submodule R₂ M₂) (h : ∀ x, f x ∈ p) : - M₁ →SL[σ₁₂] p where - cont := f.continuous.subtype_mk _ - toLinearMap := (f : M₁ →ₛₗ[σ₁₂] M₂).codRestrict p h - -@[norm_cast] -theorem coe_codRestrict (f : M₁ →SL[σ₁₂] M₂) (p : Submodule R₂ M₂) (h : ∀ x, f x ∈ p) : - (f.codRestrict p h : M₁ →ₛₗ[σ₁₂] p) = (f : M₁ →ₛₗ[σ₁₂] M₂).codRestrict p h := - rfl - -@[simp] -theorem coe_codRestrict_apply (f : M₁ →SL[σ₁₂] M₂) (p : Submodule R₂ M₂) (h : ∀ x, f x ∈ p) (x) : - (f.codRestrict p h x : M₂) = f x := - rfl - -@[simp] -theorem ker_codRestrict (f : M₁ →SL[σ₁₂] M₂) (p : Submodule R₂ M₂) (h : ∀ x, f x ∈ p) : - ker (f.codRestrict p h) = ker f := - (f : M₁ →ₛₗ[σ₁₂] M₂).ker_codRestrict p h - -/-- Restrict the codomain of a continuous linear map `f` to `f.range`. -/ -abbrev rangeRestrict [RingHomSurjective σ₁₂] (f : M₁ →SL[σ₁₂] M₂) := - f.codRestrict (LinearMap.range f) (LinearMap.mem_range_self f) - -@[simp] -theorem coe_rangeRestrict [RingHomSurjective σ₁₂] (f : M₁ →SL[σ₁₂] M₂) : - (f.rangeRestrict : M₁ →ₛₗ[σ₁₂] LinearMap.range f) = (f : M₁ →ₛₗ[σ₁₂] M₂).rangeRestrict := - rfl - -/-- `Submodule.subtype` as a `ContinuousLinearMap`. -/ -def _root_.Submodule.subtypeL (p : Submodule R₁ M₁) : p →L[R₁] M₁ where - cont := continuous_subtype_val - toLinearMap := p.subtype - -@[simp, norm_cast] -theorem _root_.Submodule.coe_subtypeL (p : Submodule R₁ M₁) : - (p.subtypeL : p →ₗ[R₁] M₁) = p.subtype := - rfl - -@[simp] -theorem _root_.Submodule.coe_subtypeL' (p : Submodule R₁ M₁) : ⇑p.subtypeL = p.subtype := - rfl - -@[simp] -- @[norm_cast] -- Porting note: A theorem with this can't have a rhs starting with `↑`. -theorem _root_.Submodule.subtypeL_apply (p : Submodule R₁ M₁) (x : p) : p.subtypeL x = x := - rfl - -@[simp] -theorem _root_.Submodule.range_subtypeL (p : Submodule R₁ M₁) : range p.subtypeL = p := - Submodule.range_subtype _ - -@[simp] -theorem _root_.Submodule.ker_subtypeL (p : Submodule R₁ M₁) : ker p.subtypeL = ⊥ := - Submodule.ker_subtype _ - -variable (R₁ M₁ M₂) - -/-- `Prod.fst` as a `ContinuousLinearMap`. -/ -def fst [Module R₁ M₂] : M₁ × M₂ →L[R₁] M₁ where - cont := continuous_fst - toLinearMap := LinearMap.fst R₁ M₁ M₂ - -/-- `Prod.snd` as a `ContinuousLinearMap`. -/ -def snd [Module R₁ M₂] : M₁ × M₂ →L[R₁] M₂ where - cont := continuous_snd - toLinearMap := LinearMap.snd R₁ M₁ M₂ - -variable {R₁ M₁ M₂} - -@[simp, norm_cast] -theorem coe_fst [Module R₁ M₂] : ↑(fst R₁ M₁ M₂) = LinearMap.fst R₁ M₁ M₂ := - rfl - -@[simp, norm_cast] -theorem coe_fst' [Module R₁ M₂] : ⇑(fst R₁ M₁ M₂) = Prod.fst := - rfl - -@[simp, norm_cast] -theorem coe_snd [Module R₁ M₂] : ↑(snd R₁ M₁ M₂) = LinearMap.snd R₁ M₁ M₂ := - rfl - -@[simp, norm_cast] -theorem coe_snd' [Module R₁ M₂] : ⇑(snd R₁ M₁ M₂) = Prod.snd := - rfl - -@[simp] -theorem fst_prod_snd [Module R₁ M₂] : (fst R₁ M₁ M₂).prod (snd R₁ M₁ M₂) = id R₁ (M₁ × M₂) := - ext fun ⟨_x, _y⟩ => rfl - -@[simp] -theorem fst_comp_prod [Module R₁ M₂] [Module R₁ M₃] (f : M₁ →L[R₁] M₂) (g : M₁ →L[R₁] M₃) : - (fst R₁ M₂ M₃).comp (f.prod g) = f := - ext fun _x => rfl - -@[simp] -theorem snd_comp_prod [Module R₁ M₂] [Module R₁ M₃] (f : M₁ →L[R₁] M₂) (g : M₁ →L[R₁] M₃) : - (snd R₁ M₂ M₃).comp (f.prod g) = g := - ext fun _x => rfl - -/-- `Prod.map` of two continuous linear maps. -/ -def prodMap [Module R₁ M₂] [Module R₁ M₃] [Module R₁ M₄] (f₁ : M₁ →L[R₁] M₂) (f₂ : M₃ →L[R₁] M₄) : - M₁ × M₃ →L[R₁] M₂ × M₄ := - (f₁.comp (fst R₁ M₁ M₃)).prod (f₂.comp (snd R₁ M₁ M₃)) - -@[simp, norm_cast] -theorem coe_prodMap [Module R₁ M₂] [Module R₁ M₃] [Module R₁ M₄] (f₁ : M₁ →L[R₁] M₂) - (f₂ : M₃ →L[R₁] M₄) : ↑(f₁.prodMap f₂) = (f₁ : M₁ →ₗ[R₁] M₂).prodMap (f₂ : M₃ →ₗ[R₁] M₄) := - rfl - -@[simp, norm_cast] -theorem coe_prodMap' [Module R₁ M₂] [Module R₁ M₃] [Module R₁ M₄] (f₁ : M₁ →L[R₁] M₂) - (f₂ : M₃ →L[R₁] M₄) : ⇑(f₁.prodMap f₂) = Prod.map f₁ f₂ := - rfl - -/-- The continuous linear map given by `(x, y) ↦ f₁ x + f₂ y`. -/ -def coprod [Module R₁ M₂] [Module R₁ M₃] [ContinuousAdd M₃] (f₁ : M₁ →L[R₁] M₃) - (f₂ : M₂ →L[R₁] M₃) : M₁ × M₂ →L[R₁] M₃ := - ⟨LinearMap.coprod f₁ f₂, (f₁.cont.comp continuous_fst).add (f₂.cont.comp continuous_snd)⟩ - -@[norm_cast, simp] -theorem coe_coprod [Module R₁ M₂] [Module R₁ M₃] [ContinuousAdd M₃] (f₁ : M₁ →L[R₁] M₃) - (f₂ : M₂ →L[R₁] M₃) : (f₁.coprod f₂ : M₁ × M₂ →ₗ[R₁] M₃) = LinearMap.coprod f₁ f₂ := - rfl - -@[simp] -theorem coprod_apply [Module R₁ M₂] [Module R₁ M₃] [ContinuousAdd M₃] (f₁ : M₁ →L[R₁] M₃) - (f₂ : M₂ →L[R₁] M₃) (x) : f₁.coprod f₂ x = f₁ x.1 + f₂ x.2 := - rfl - -theorem range_coprod [Module R₁ M₂] [Module R₁ M₃] [ContinuousAdd M₃] (f₁ : M₁ →L[R₁] M₃) - (f₂ : M₂ →L[R₁] M₃) : range (f₁.coprod f₂) = range f₁ ⊔ range f₂ := - LinearMap.range_coprod _ _ - -theorem comp_fst_add_comp_snd [Module R₁ M₂] [Module R₁ M₃] [ContinuousAdd M₃] (f : M₁ →L[R₁] M₃) - (g : M₂ →L[R₁] M₃) : - f.comp (ContinuousLinearMap.fst R₁ M₁ M₂) + g.comp (ContinuousLinearMap.snd R₁ M₁ M₂) = - f.coprod g := - rfl +section Quotient -theorem coprod_inl_inr [ContinuousAdd M₁] [ContinuousAdd M'₁] : - (ContinuousLinearMap.inl R₁ M₁ M'₁).coprod (ContinuousLinearMap.inr R₁ M₁ M'₁) = - ContinuousLinearMap.id R₁ (M₁ × M'₁) := by - apply coe_injective; apply LinearMap.coprod_inl_inr +namespace Submodule -section +variable {R M : Type*} [Ring R] [AddCommGroup M] [Module R M] [TopologicalSpace M] + (S : Submodule R M) -variable {R S : Type*} [Semiring R] [Semiring S] [Module R M₁] [Module R M₂] [Module R S] - [Module S M₂] [IsScalarTower R S M₂] [TopologicalSpace S] [ContinuousSMul S M₂] +-- Porting note: This is required in Lean4. +instance _root_.QuotientModule.Quotient.topologicalSpace : TopologicalSpace (M ⧸ S) := + inferInstanceAs (TopologicalSpace (Quotient S.quotientRel)) -/-- The linear map `fun x => c x • f`. Associates to a scalar-valued linear map and an element of -`M₂` the `M₂`-valued linear map obtained by multiplying the two (a.k.a. tensoring by `M₂`). -See also `ContinuousLinearMap.smulRightₗ` and `ContinuousLinearMap.smulRightL`. -/ -def smulRight (c : M₁ →L[R] S) (f : M₂) : M₁ →L[R] M₂ := - { c.toLinearMap.smulRight f with cont := c.2.smul continuous_const } +theorem isOpenMap_mkQ [ContinuousAdd M] : IsOpenMap S.mkQ := + QuotientAddGroup.isOpenMap_coe -@[simp] -theorem smulRight_apply {c : M₁ →L[R] S} {f : M₂} {x : M₁} : - (smulRight c f : M₁ → M₂) x = c x • f := - rfl +theorem isOpenQuotientMap_mkQ [ContinuousAdd M] : IsOpenQuotientMap S.mkQ := + QuotientAddGroup.isOpenQuotientMap_mk -end +instance topologicalAddGroup_quotient [TopologicalAddGroup M] : TopologicalAddGroup (M ⧸ S) := + inferInstanceAs <| TopologicalAddGroup (M ⧸ S.toAddSubgroup) -variable [Module R₁ M₂] [TopologicalSpace R₁] [ContinuousSMul R₁ M₂] +instance continuousSMul_quotient [TopologicalSpace R] [TopologicalAddGroup M] [ContinuousSMul R M] : + 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 -@[simp] -theorem smulRight_one_one (c : R₁ →L[R₁] M₂) : smulRight (1 : R₁ →L[R₁] R₁) (c 1) = c := by - ext - simp [← ContinuousLinearMap.map_smul_of_tower] +instance t3_quotient_of_isClosed [TopologicalAddGroup M] [IsClosed (S : Set M)] : + T3Space (M ⧸ S) := + letI : IsClosed (S.toAddSubgroup : Set M) := ‹_› + QuotientAddGroup.instT3Space S.toAddSubgroup -@[simp] -theorem smulRight_one_eq_iff {f f' : M₂} : - smulRight (1 : R₁ →L[R₁] R₁) f = smulRight (1 : R₁ →L[R₁] R₁) f' ↔ f = f' := by - simp only [ContinuousLinearMap.ext_ring_iff, smulRight_apply, one_apply, one_smul] - -theorem smulRight_comp [ContinuousMul R₁] {x : M₂} {c : R₁} : - (smulRight (1 : R₁ →L[R₁] R₁) x).comp (smulRight (1 : R₁ →L[R₁] R₁) c) = - smulRight (1 : R₁ →L[R₁] R₁) (c • x) := by - ext - simp [mul_smul] - -section ToSpanSingleton - -variable (R₁) -variable [ContinuousSMul R₁ M₁] - -/-- Given an element `x` of a topological space `M` over a semiring `R`, the natural continuous -linear map from `R` to `M` by taking multiples of `x`. -/ -def toSpanSingleton (x : M₁) : R₁ →L[R₁] M₁ where - toLinearMap := LinearMap.toSpanSingleton R₁ M₁ x - cont := continuous_id.smul continuous_const - -theorem toSpanSingleton_apply (x : M₁) (r : R₁) : toSpanSingleton R₁ x r = r • x := - rfl - -theorem toSpanSingleton_add [ContinuousAdd M₁] (x y : M₁) : - toSpanSingleton R₁ (x + y) = toSpanSingleton R₁ x + toSpanSingleton R₁ y := by - ext1; simp [toSpanSingleton_apply] - -theorem toSpanSingleton_smul' {α} [Monoid α] [DistribMulAction α M₁] [ContinuousConstSMul α M₁] - [SMulCommClass R₁ α M₁] (c : α) (x : M₁) : - toSpanSingleton R₁ (c • x) = c • toSpanSingleton R₁ x := by - ext1; rw [toSpanSingleton_apply, smul_apply, toSpanSingleton_apply, smul_comm] - -/-- A special case of `to_span_singleton_smul'` for when `R` is commutative. -/ -theorem toSpanSingleton_smul (R) {M₁} [CommSemiring R] [AddCommMonoid M₁] [Module R M₁] - [TopologicalSpace R] [TopologicalSpace M₁] [ContinuousSMul R M₁] (c : R) (x : M₁) : - toSpanSingleton R (c • x) = c • toSpanSingleton R x := - toSpanSingleton_smul' R c x - -end ToSpanSingleton - -end Semiring - -section Pi - -variable {R : Type*} [Semiring R] {M : Type*} [TopologicalSpace M] [AddCommMonoid M] [Module R M] - {M₂ : Type*} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R M₂] {ι : Type*} {φ : ι → Type*} - [∀ i, TopologicalSpace (φ i)] [∀ i, AddCommMonoid (φ i)] [∀ i, Module R (φ i)] - -/-- `pi` construction for continuous linear functions. From a family of continuous linear functions -it produces a continuous linear function into a family of topological modules. -/ -def pi (f : ∀ i, M →L[R] φ i) : M →L[R] ∀ i, φ i := - ⟨LinearMap.pi fun i => f i, continuous_pi fun i => (f i).continuous⟩ - -@[simp] -theorem coe_pi' (f : ∀ i, M →L[R] φ i) : ⇑(pi f) = fun c i => f i c := - rfl - -@[simp] -theorem coe_pi (f : ∀ i, M →L[R] φ i) : (pi f : M →ₗ[R] ∀ i, φ i) = LinearMap.pi fun i => f i := - rfl - -theorem pi_apply (f : ∀ i, M →L[R] φ i) (c : M) (i : ι) : pi f c i = f i c := - rfl - -theorem pi_eq_zero (f : ∀ i, M →L[R] φ i) : pi f = 0 ↔ ∀ i, f i = 0 := by - simp only [ContinuousLinearMap.ext_iff, pi_apply, funext_iff] - exact forall_swap - -theorem pi_zero : pi (fun _ => 0 : ∀ i, M →L[R] φ i) = 0 := - ext fun _ => rfl - -theorem pi_comp (f : ∀ i, M →L[R] φ i) (g : M₂ →L[R] M) : - (pi f).comp g = pi fun i => (f i).comp g := - rfl - -/-- The projections from a family of topological modules are continuous linear maps. -/ -def proj (i : ι) : (∀ i, φ i) →L[R] φ i := - ⟨LinearMap.proj i, continuous_apply _⟩ - -@[simp] -theorem proj_apply (i : ι) (b : ∀ i, φ i) : (proj i : (∀ i, φ i) →L[R] φ i) b = b i := - rfl - -theorem proj_pi (f : ∀ i, M₂ →L[R] φ i) (i : ι) : (proj i).comp (pi f) = f i := - ext fun _c => rfl - -theorem iInf_ker_proj : (⨅ i, ker (proj i : (∀ i, φ i) →L[R] φ i) : Submodule R (∀ i, φ i)) = ⊥ := - LinearMap.iInf_ker_proj - -variable (R φ) - -/-- Given a function `f : α → ι`, it induces a continuous linear function by right composition on -product types. For `f = Subtype.val`, this corresponds to forgetting some set of variables. -/ -def _root_.Pi.compRightL {α : Type*} (f : α → ι) : ((i : ι) → φ i) →L[R] ((i : α) → φ (f i)) where - toFun := fun v i ↦ v (f i) - map_add' := by intros; ext; simp - map_smul' := by intros; ext; simp - cont := by continuity - -@[simp] lemma _root_.Pi.compRightL_apply {α : Type*} (f : α → ι) (v : (i : ι) → φ i) (i : α) : - Pi.compRightL R φ f v i = v (f i) := rfl - -/-- If `I` and `J` are complementary index sets, the product of the kernels of the `J`th projections -of `φ` is linearly equivalent to the product over `I`. -/ -def iInfKerProjEquiv {I J : Set ι} [DecidablePred fun i => i ∈ I] (hd : Disjoint I J) - (hu : Set.univ ⊆ I ∪ J) : - (⨅ i ∈ J, ker (proj i : (∀ i, φ i) →L[R] φ i) : - Submodule R (∀ i, φ i)) ≃L[R] ∀ i : I, φ i where - toLinearEquiv := LinearMap.iInfKerProjEquiv R φ hd hu - continuous_toFun := - continuous_pi fun i => - Continuous.comp (continuous_apply (π := φ) i) <| - @continuous_subtype_val _ _ fun x => - x ∈ (⨅ i ∈ J, ker (proj i : (∀ i, φ i) →L[R] φ i) : Submodule R (∀ i, φ i)) - continuous_invFun := - Continuous.subtype_mk - (continuous_pi fun i => by - dsimp - split_ifs <;> [apply continuous_apply; exact continuous_zero]) - _ - -end Pi - -section Ring - -variable {R : Type*} [Ring R] {R₂ : Type*} [Ring R₂] {R₃ : Type*} [Ring R₃] {M : Type*} - [TopologicalSpace M] [AddCommGroup M] {M₂ : Type*} [TopologicalSpace M₂] [AddCommGroup M₂] - {M₃ : Type*} [TopologicalSpace M₃] [AddCommGroup M₃] {M₄ : Type*} [TopologicalSpace M₄] - [AddCommGroup M₄] [Module R M] [Module R₂ M₂] [Module R₃ M₃] {σ₁₂ : R →+* R₂} {σ₂₃ : R₂ →+* R₃} - {σ₁₃ : R →+* R₃} - -section - -protected theorem map_neg (f : M →SL[σ₁₂] M₂) (x : M) : f (-x) = -f x := by - exact map_neg f x - -protected theorem map_sub (f : M →SL[σ₁₂] M₂) (x y : M) : f (x - y) = f x - f y := by - exact map_sub f x y - -@[simp] -theorem sub_apply' (f g : M →SL[σ₁₂] M₂) (x : M) : ((f : M →ₛₗ[σ₁₂] M₂) - g) x = f x - g x := - rfl - -end - -section - -variable [Module R M₂] [Module R M₃] [Module R M₄] - -theorem range_prod_eq {f : M →L[R] M₂} {g : M →L[R] M₃} (h : ker f ⊔ ker g = ⊤) : - range (f.prod g) = (range f).prod (range g) := - LinearMap.range_prod_eq h - -theorem ker_prod_ker_le_ker_coprod [ContinuousAdd M₃] (f : M →L[R] M₃) (g : M₂ →L[R] M₃) : - (LinearMap.ker f).prod (LinearMap.ker g) ≤ LinearMap.ker (f.coprod g) := - LinearMap.ker_prod_ker_le_ker_coprod f.toLinearMap g.toLinearMap - -theorem ker_coprod_of_disjoint_range [ContinuousAdd M₃] (f : M →L[R] M₃) (g : M₂ →L[R] M₃) - (hd : Disjoint (range f) (range g)) : - LinearMap.ker (f.coprod g) = (LinearMap.ker f).prod (LinearMap.ker g) := - LinearMap.ker_coprod_of_disjoint_range f.toLinearMap g.toLinearMap hd - -end - -section - -variable [TopologicalAddGroup M₂] - -instance neg : Neg (M →SL[σ₁₂] M₂) := - ⟨fun f => ⟨-f, f.2.neg⟩⟩ - -@[simp] -theorem neg_apply (f : M →SL[σ₁₂] M₂) (x : M) : (-f) x = -f x := - rfl - -@[simp, norm_cast] -theorem coe_neg (f : M →SL[σ₁₂] M₂) : (↑(-f) : M →ₛₗ[σ₁₂] M₂) = -f := - rfl - -@[norm_cast] -theorem coe_neg' (f : M →SL[σ₁₂] M₂) : ⇑(-f) = -f := - rfl - -instance sub : Sub (M →SL[σ₁₂] M₂) := - ⟨fun f g => ⟨f - g, f.2.sub g.2⟩⟩ - -instance addCommGroup : AddCommGroup (M →SL[σ₁₂] M₂) where - __ := ContinuousLinearMap.addCommMonoid - neg := (-·) - sub := (· - ·) - sub_eq_add_neg _ _ := by ext; apply sub_eq_add_neg - nsmul := (· • ·) - zsmul := (· • ·) - zsmul_zero' f := by ext; simp - zsmul_succ' n f := by ext; simp [add_smul, add_comm] - zsmul_neg' n f := by ext; simp [add_smul] - neg_add_cancel _ := by ext; apply neg_add_cancel - -theorem sub_apply (f g : M →SL[σ₁₂] M₂) (x : M) : (f - g) x = f x - g x := - rfl - -@[simp, norm_cast] -theorem coe_sub (f g : M →SL[σ₁₂] M₂) : (↑(f - g) : M →ₛₗ[σ₁₂] M₂) = f - g := - rfl - -@[simp, norm_cast] -theorem coe_sub' (f g : M →SL[σ₁₂] M₂) : ⇑(f - g) = f - g := - rfl - -end - -@[simp] -theorem comp_neg [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] [TopologicalAddGroup M₂] [TopologicalAddGroup M₃] - (g : M₂ →SL[σ₂₃] M₃) (f : M →SL[σ₁₂] M₂) : g.comp (-f) = -g.comp f := by - ext x - simp - -@[simp] -theorem neg_comp [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] [TopologicalAddGroup M₃] (g : M₂ →SL[σ₂₃] M₃) - (f : M →SL[σ₁₂] M₂) : (-g).comp f = -g.comp f := by - ext - simp - -@[simp] -theorem comp_sub [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] [TopologicalAddGroup M₂] [TopologicalAddGroup M₃] - (g : M₂ →SL[σ₂₃] M₃) (f₁ f₂ : M →SL[σ₁₂] M₂) : g.comp (f₁ - f₂) = g.comp f₁ - g.comp f₂ := by - ext - simp - -@[simp] -theorem sub_comp [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] [TopologicalAddGroup M₃] (g₁ g₂ : M₂ →SL[σ₂₃] M₃) - (f : M →SL[σ₁₂] M₂) : (g₁ - g₂).comp f = g₁.comp f - g₂.comp f := by - ext - simp - -instance ring [TopologicalAddGroup M] : Ring (M →L[R] M) where - __ := ContinuousLinearMap.semiring - __ := ContinuousLinearMap.addCommGroup - intCast z := z • (1 : M →L[R] M) - intCast_ofNat := natCast_zsmul _ - intCast_negSucc := negSucc_zsmul _ - -@[simp] -theorem intCast_apply [TopologicalAddGroup M] (z : ℤ) (m : M) : (↑z : M →L[R] M) m = z • m := - rfl - -theorem smulRight_one_pow [TopologicalSpace R] [TopologicalRing R] (c : R) (n : ℕ) : - smulRight (1 : R →L[R] R) c ^ n = smulRight (1 : R →L[R] R) (c ^ n) := by - induction n with - | zero => ext; simp - | succ n ihn => rw [pow_succ, ihn, mul_def, smulRight_comp, smul_eq_mul, pow_succ'] - -section - -variable {σ₂₁ : R₂ →+* R} [RingHomInvPair σ₁₂ σ₂₁] - - -/-- Given a right inverse `f₂ : M₂ →L[R] M` to `f₁ : M →L[R] M₂`, -`projKerOfRightInverse f₁ f₂ h` is the projection `M →L[R] LinearMap.ker f₁` along -`LinearMap.range f₂`. -/ -def projKerOfRightInverse [TopologicalAddGroup M] (f₁ : M →SL[σ₁₂] M₂) (f₂ : M₂ →SL[σ₂₁] M) - (h : Function.RightInverse f₂ f₁) : M →L[R] LinearMap.ker f₁ := - (id R M - f₂.comp f₁).codRestrict (LinearMap.ker f₁) fun x => by simp [h (f₁ x)] - -@[simp] -theorem coe_projKerOfRightInverse_apply [TopologicalAddGroup M] (f₁ : M →SL[σ₁₂] M₂) - (f₂ : M₂ →SL[σ₂₁] M) (h : Function.RightInverse f₂ f₁) (x : M) : - (f₁.projKerOfRightInverse f₂ h x : M) = x - f₂ (f₁ x) := - rfl - -@[simp] -theorem projKerOfRightInverse_apply_idem [TopologicalAddGroup M] (f₁ : M →SL[σ₁₂] M₂) - (f₂ : M₂ →SL[σ₂₁] M) (h : Function.RightInverse f₂ f₁) (x : LinearMap.ker f₁) : - f₁.projKerOfRightInverse f₂ h x = x := by - ext1 - simp - -@[simp] -theorem projKerOfRightInverse_comp_inv [TopologicalAddGroup M] (f₁ : M →SL[σ₁₂] M₂) - (f₂ : M₂ →SL[σ₂₁] M) (h : Function.RightInverse f₂ f₁) (y : M₂) : - f₁.projKerOfRightInverse f₂ h (f₂ y) = 0 := - Subtype.ext_iff_val.2 <| by simp [h y] - -end - -end Ring - -section DivisionMonoid - -variable {R M : Type*} - -/-- A nonzero continuous linear functional is open. -/ -protected theorem isOpenMap_of_ne_zero [TopologicalSpace R] [DivisionRing R] [ContinuousSub R] - [AddCommGroup M] [TopologicalSpace M] [ContinuousAdd M] [Module R M] [ContinuousSMul R M] - (f : M →L[R] R) (hf : f ≠ 0) : IsOpenMap f := - let ⟨x, hx⟩ := exists_ne_zero hf - IsOpenMap.of_sections fun y => - ⟨fun a => y + (a - f y) • (f x)⁻¹ • x, Continuous.continuousAt <| by continuity, by simp, - fun a => by simp [hx]⟩ - -end DivisionMonoid - -section SMulMonoid - --- The M's are used for semilinear maps, and the N's for plain linear maps -variable {R R₂ R₃ S S₃ : Type*} [Semiring R] [Semiring R₂] [Semiring R₃] [Monoid S] [Monoid S₃] - {M : Type*} [TopologicalSpace M] [AddCommMonoid M] [Module R M] {M₂ : Type*} - [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₂ M₂] {M₃ : Type*} [TopologicalSpace M₃] - [AddCommMonoid M₃] [Module R₃ M₃] {N₂ : Type*} [TopologicalSpace N₂] [AddCommMonoid N₂] - [Module R N₂] {N₃ : Type*} [TopologicalSpace N₃] [AddCommMonoid N₃] [Module R N₃] - [DistribMulAction S₃ M₃] [SMulCommClass R₃ S₃ M₃] [ContinuousConstSMul S₃ M₃] - [DistribMulAction S N₃] [SMulCommClass R S N₃] [ContinuousConstSMul S N₃] {σ₁₂ : R →+* R₂} - {σ₂₃ : R₂ →+* R₃} {σ₁₃ : R →+* R₃} [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] - -@[simp] -theorem smul_comp (c : S₃) (h : M₂ →SL[σ₂₃] M₃) (f : M →SL[σ₁₂] M₂) : - (c • h).comp f = c • h.comp f := - rfl - -variable [DistribMulAction S₃ M₂] [ContinuousConstSMul S₃ M₂] [SMulCommClass R₂ S₃ M₂] -variable [DistribMulAction S N₂] [ContinuousConstSMul S N₂] [SMulCommClass R S N₂] - -@[simp] -theorem comp_smul [LinearMap.CompatibleSMul N₂ N₃ S R] (hₗ : N₂ →L[R] N₃) (c : S) - (fₗ : M →L[R] N₂) : hₗ.comp (c • fₗ) = c • hₗ.comp fₗ := by - ext x - exact hₗ.map_smul_of_tower c (fₗ x) - -@[simp] -theorem comp_smulₛₗ [SMulCommClass R₂ R₂ M₂] [SMulCommClass R₃ R₃ M₃] [ContinuousConstSMul R₂ M₂] - [ContinuousConstSMul R₃ M₃] (h : M₂ →SL[σ₂₃] M₃) (c : R₂) (f : M →SL[σ₁₂] M₂) : - h.comp (c • f) = σ₂₃ c • h.comp f := by - ext x - simp only [coe_smul', coe_comp', Function.comp_apply, Pi.smul_apply, - ContinuousLinearMap.map_smulₛₗ] - -instance distribMulAction [ContinuousAdd M₂] : DistribMulAction S₃ (M →SL[σ₁₂] M₂) where - smul_add a f g := ext fun x => smul_add a (f x) (g x) - smul_zero a := ext fun _ => smul_zero a - -end SMulMonoid - -section SMul - --- The M's are used for semilinear maps, and the N's for plain linear maps -variable {R R₂ R₃ S S₃ : Type*} [Semiring R] [Semiring R₂] [Semiring R₃] [Semiring S] [Semiring S₃] - {M : Type*} [TopologicalSpace M] [AddCommMonoid M] [Module R M] {M₂ : Type*} - [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₂ M₂] {M₃ : Type*} [TopologicalSpace M₃] - [AddCommMonoid M₃] [Module R₃ M₃] {N₂ : Type*} [TopologicalSpace N₂] [AddCommMonoid N₂] - [Module R N₂] {N₃ : Type*} [TopologicalSpace N₃] [AddCommMonoid N₃] [Module R N₃] [Module S₃ M₃] - [SMulCommClass R₃ S₃ M₃] [ContinuousConstSMul S₃ M₃] [Module S N₂] [ContinuousConstSMul S N₂] - [SMulCommClass R S N₂] [Module S N₃] [SMulCommClass R S N₃] [ContinuousConstSMul S N₃] - {σ₁₂ : R →+* R₂} {σ₂₃ : R₂ →+* R₃} {σ₁₃ : R →+* R₃} [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] (c : S) - (h : M₂ →SL[σ₂₃] M₃) (f : M →SL[σ₁₂] M₂) - -/-- `ContinuousLinearMap.prod` as an `Equiv`. -/ -@[simps apply] -def prodEquiv : (M →L[R] N₂) × (M →L[R] N₃) ≃ (M →L[R] N₂ × N₃) where - toFun f := f.1.prod f.2 - invFun f := ⟨(fst _ _ _).comp f, (snd _ _ _).comp f⟩ - left_inv f := by ext <;> rfl - right_inv f := by ext <;> rfl - -theorem prod_ext_iff {f g : M × N₂ →L[R] N₃} : - f = g ↔ f.comp (inl _ _ _) = g.comp (inl _ _ _) ∧ f.comp (inr _ _ _) = g.comp (inr _ _ _) := by - simp only [← coe_inj, LinearMap.prod_ext_iff] - rfl - -@[ext] -theorem prod_ext {f g : M × N₂ →L[R] N₃} (hl : f.comp (inl _ _ _) = g.comp (inl _ _ _)) - (hr : f.comp (inr _ _ _) = g.comp (inr _ _ _)) : f = g := - prod_ext_iff.2 ⟨hl, hr⟩ - -variable [ContinuousAdd M₂] [ContinuousAdd M₃] [ContinuousAdd N₂] - -instance module : Module S₃ (M →SL[σ₁₃] M₃) where - zero_smul _ := ext fun _ => zero_smul S₃ _ - add_smul _ _ _ := ext fun _ => add_smul _ _ _ - -instance isCentralScalar [Module S₃ᵐᵒᵖ M₃] [IsCentralScalar S₃ M₃] : - IsCentralScalar S₃ (M →SL[σ₁₃] M₃) where - op_smul_eq_smul _ _ := ext fun _ => op_smul_eq_smul _ _ - -variable (S) [ContinuousAdd N₃] - -/-- `ContinuousLinearMap.prod` as a `LinearEquiv`. -/ -@[simps apply] -def prodₗ : ((M →L[R] N₂) × (M →L[R] N₃)) ≃ₗ[S] M →L[R] N₂ × N₃ := - { prodEquiv with - map_add' := fun _f _g => rfl - map_smul' := fun _c _f => rfl } - -/-- The coercion from `M →L[R] M₂` to `M →ₗ[R] M₂`, as a linear map. -/ -@[simps] -def coeLM : (M →L[R] N₃) →ₗ[S] M →ₗ[R] N₃ where - toFun := (↑) - map_add' f g := coe_add f g - map_smul' c f := coe_smul c f - -variable {S} (σ₁₃) - -/-- The coercion from `M →SL[σ] M₂` to `M →ₛₗ[σ] M₂`, as a linear map. -/ -@[simps] -def coeLMₛₗ : (M →SL[σ₁₃] M₃) →ₗ[S₃] M →ₛₗ[σ₁₃] M₃ where - toFun := (↑) - map_add' f g := coe_add f g - map_smul' c f := coe_smul c f - -end SMul - -section SMulRightₗ - -variable {R S T M M₂ : Type*} [Semiring R] [Semiring S] [Semiring T] [Module R S] - [AddCommMonoid M₂] [Module R M₂] [Module S M₂] [IsScalarTower R S M₂] [TopologicalSpace S] - [TopologicalSpace M₂] [ContinuousSMul S M₂] [TopologicalSpace M] [AddCommMonoid M] [Module R M] - [ContinuousAdd M₂] [Module T M₂] [ContinuousConstSMul T M₂] [SMulCommClass R T M₂] - [SMulCommClass S T M₂] - -/-- Given `c : E →L[R] S`, `c.smulRightₗ` is the linear map from `F` to `E →L[R] F` -sending `f` to `fun e => c e • f`. See also `ContinuousLinearMap.smulRightL`. -/ -def smulRightₗ (c : M →L[R] S) : M₂ →ₗ[T] M →L[R] M₂ where - toFun := c.smulRight - map_add' x y := by - ext e - apply smul_add (c e) - map_smul' a x := by - ext e - dsimp - apply smul_comm - -@[simp] -theorem coe_smulRightₗ (c : M →L[R] S) : ⇑(smulRightₗ c : M₂ →ₗ[T] M →L[R] M₂) = c.smulRight := - rfl - -end SMulRightₗ - -section CommRing - -variable {R : Type*} [CommRing R] {M : Type*} [TopologicalSpace M] [AddCommGroup M] {M₂ : Type*} - [TopologicalSpace M₂] [AddCommGroup M₂] {M₃ : Type*} [TopologicalSpace M₃] [AddCommGroup M₃] - [Module R M] [Module R M₂] [Module R M₃] - -variable [TopologicalAddGroup M₂] [ContinuousConstSMul R M₂] - -instance algebra : Algebra R (M₂ →L[R] M₂) := - Algebra.ofModule smul_comp fun _ _ _ => comp_smul _ _ _ - -@[simp] theorem algebraMap_apply (r : R) (m : M₂) : algebraMap R (M₂ →L[R] M₂) r m = r • m := rfl - -end CommRing - -section RestrictScalars - -variable {A M M₂ : Type*} [Ring A] [AddCommGroup M] [AddCommGroup M₂] [Module A M] [Module A M₂] - [TopologicalSpace M] [TopologicalSpace M₂] (R : Type*) [Ring R] [Module R M] [Module R M₂] - [LinearMap.CompatibleSMul M M₂ R A] - -/-- If `A` is an `R`-algebra, then a continuous `A`-linear map can be interpreted as a continuous -`R`-linear map. We assume `LinearMap.CompatibleSMul M M₂ R A` to match assumptions of -`LinearMap.map_smul_of_tower`. -/ -def restrictScalars (f : M →L[A] M₂) : M →L[R] M₂ := - ⟨(f : M →ₗ[A] M₂).restrictScalars R, f.continuous⟩ - -variable {R} - -@[simp] -- @[norm_cast] -- Porting note: This theorem can't be a `norm_cast` theorem. -theorem coe_restrictScalars (f : M →L[A] M₂) : - (f.restrictScalars R : M →ₗ[R] M₂) = (f : M →ₗ[A] M₂).restrictScalars R := - rfl - -@[simp] -theorem coe_restrictScalars' (f : M →L[A] M₂) : ⇑(f.restrictScalars R) = f := - rfl - -@[simp] -theorem restrictScalars_zero : (0 : M →L[A] M₂).restrictScalars R = 0 := - rfl - -section - -variable [TopologicalAddGroup M₂] - -@[simp] -theorem restrictScalars_add (f g : M →L[A] M₂) : - (f + g).restrictScalars R = f.restrictScalars R + g.restrictScalars R := - rfl - -@[simp] -theorem restrictScalars_neg (f : M →L[A] M₂) : (-f).restrictScalars R = -f.restrictScalars R := - rfl - -end - -variable {S : Type*} -variable [Ring S] [Module S M₂] [ContinuousConstSMul S M₂] [SMulCommClass A S M₂] - [SMulCommClass R S M₂] - -@[simp] -theorem restrictScalars_smul (c : S) (f : M →L[A] M₂) : - (c • f).restrictScalars R = c • f.restrictScalars R := - rfl - -variable (A M M₂ R S) -variable [TopologicalAddGroup M₂] - -/-- `ContinuousLinearMap.restrictScalars` as a `LinearMap`. See also -`ContinuousLinearMap.restrictScalarsL`. -/ -def restrictScalarsₗ : (M →L[A] M₂) →ₗ[S] M →L[R] M₂ where - toFun := restrictScalars R - map_add' := restrictScalars_add - map_smul' := restrictScalars_smul - -variable {A M M₂ R S} - -@[simp] -theorem coe_restrictScalarsₗ : ⇑(restrictScalarsₗ A M M₂ R S) = restrictScalars R := - rfl - -end RestrictScalars - -end ContinuousLinearMap - -namespace ContinuousLinearEquiv - -section AddCommMonoid - -variable {R₁ : Type*} {R₂ : Type*} {R₃ : Type*} [Semiring R₁] [Semiring R₂] [Semiring R₃] - {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] - {σ₂₃ : R₂ →+* R₃} {σ₃₂ : R₃ →+* R₂} [RingHomInvPair σ₂₃ σ₃₂] [RingHomInvPair σ₃₂ σ₂₃] - {σ₁₃ : R₁ →+* R₃} {σ₃₁ : R₃ →+* R₁} [RingHomInvPair σ₁₃ σ₃₁] [RingHomInvPair σ₃₁ σ₁₃] - [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] [RingHomCompTriple σ₃₂ σ₂₁ σ₃₁] {M₁ : Type*} - [TopologicalSpace M₁] [AddCommMonoid M₁] - {M₂ : Type*} [TopologicalSpace M₂] [AddCommMonoid M₂] {M₃ : Type*} [TopologicalSpace M₃] - [AddCommMonoid M₃] {M₄ : Type*} [TopologicalSpace M₄] [AddCommMonoid M₄] [Module R₁ M₁] - [Module R₂ M₂] [Module R₃ M₃] - -/-- A continuous linear equivalence induces a continuous linear map. -/ -@[coe] -def toContinuousLinearMap (e : M₁ ≃SL[σ₁₂] M₂) : M₁ →SL[σ₁₂] M₂ := - { e.toLinearEquiv.toLinearMap with cont := e.continuous_toFun } - -/-- Coerce continuous linear equivs to continuous linear maps. -/ -instance ContinuousLinearMap.coe : Coe (M₁ ≃SL[σ₁₂] M₂) (M₁ →SL[σ₁₂] M₂) := - ⟨toContinuousLinearMap⟩ - -instance equivLike : - EquivLike (M₁ ≃SL[σ₁₂] M₂) M₁ M₂ where - coe f := f.toFun - inv f := f.invFun - coe_injective' f g h₁ h₂ := by - cases' f with f' _ - cases' g with g' _ - rcases f' with ⟨⟨⟨_, _⟩, _⟩, _⟩ - rcases g' with ⟨⟨⟨_, _⟩, _⟩, _⟩ - congr - left_inv f := f.left_inv - right_inv f := f.right_inv - -instance continuousSemilinearEquivClass : - ContinuousSemilinearEquivClass (M₁ ≃SL[σ₁₂] M₂) σ₁₂ M₁ M₂ where - map_add f := f.map_add' - map_smulₛₗ f := f.map_smul' - map_continuous := continuous_toFun - inv_continuous := continuous_invFun - -theorem coe_apply (e : M₁ ≃SL[σ₁₂] M₂) (b : M₁) : (e : M₁ →SL[σ₁₂] M₂) b = e b := - rfl - -@[simp] -theorem coe_toLinearEquiv (f : M₁ ≃SL[σ₁₂] M₂) : ⇑f.toLinearEquiv = f := - rfl - -@[simp, norm_cast] -theorem coe_coe (e : M₁ ≃SL[σ₁₂] M₂) : ⇑(e : M₁ →SL[σ₁₂] M₂) = e := - rfl - -theorem toLinearEquiv_injective : - Function.Injective (toLinearEquiv : (M₁ ≃SL[σ₁₂] M₂) → M₁ ≃ₛₗ[σ₁₂] M₂) := by - rintro ⟨e, _, _⟩ ⟨e', _, _⟩ rfl - rfl - -@[ext] -theorem ext {f g : M₁ ≃SL[σ₁₂] M₂} (h : (f : M₁ → M₂) = g) : f = g := - toLinearEquiv_injective <| LinearEquiv.ext <| congr_fun h - -theorem coe_injective : Function.Injective ((↑) : (M₁ ≃SL[σ₁₂] M₂) → M₁ →SL[σ₁₂] M₂) := - fun _e _e' h => ext <| funext <| ContinuousLinearMap.ext_iff.1 h - -@[simp, norm_cast] -theorem coe_inj {e e' : M₁ ≃SL[σ₁₂] M₂} : (e : M₁ →SL[σ₁₂] M₂) = e' ↔ e = e' := - coe_injective.eq_iff - -/-- A continuous linear equivalence induces a homeomorphism. -/ -def toHomeomorph (e : M₁ ≃SL[σ₁₂] M₂) : M₁ ≃ₜ M₂ := - { e with toEquiv := e.toLinearEquiv.toEquiv } - -@[simp] -theorem coe_toHomeomorph (e : M₁ ≃SL[σ₁₂] M₂) : ⇑e.toHomeomorph = e := - rfl - -theorem isOpenMap (e : M₁ ≃SL[σ₁₂] M₂) : IsOpenMap e := - (ContinuousLinearEquiv.toHomeomorph e).isOpenMap - -theorem image_closure (e : M₁ ≃SL[σ₁₂] M₂) (s : Set M₁) : e '' closure s = closure (e '' s) := - e.toHomeomorph.image_closure s - -theorem preimage_closure (e : M₁ ≃SL[σ₁₂] M₂) (s : Set M₂) : e ⁻¹' closure s = closure (e ⁻¹' s) := - e.toHomeomorph.preimage_closure s - -@[simp] -theorem isClosed_image (e : M₁ ≃SL[σ₁₂] M₂) {s : Set M₁} : IsClosed (e '' s) ↔ IsClosed s := - e.toHomeomorph.isClosed_image - -theorem map_nhds_eq (e : M₁ ≃SL[σ₁₂] M₂) (x : M₁) : map e (𝓝 x) = 𝓝 (e x) := - e.toHomeomorph.map_nhds_eq x - --- Make some straightforward lemmas available to `simp`. -theorem map_zero (e : M₁ ≃SL[σ₁₂] M₂) : e (0 : M₁) = 0 := - (e : M₁ →SL[σ₁₂] M₂).map_zero - -theorem map_add (e : M₁ ≃SL[σ₁₂] M₂) (x y : M₁) : e (x + y) = e x + e y := - (e : M₁ →SL[σ₁₂] M₂).map_add x y - -@[simp] -theorem map_smulₛₗ (e : M₁ ≃SL[σ₁₂] M₂) (c : R₁) (x : M₁) : e (c • x) = σ₁₂ c • e x := - (e : M₁ →SL[σ₁₂] M₂).map_smulₛₗ c x - -theorem map_smul [Module R₁ M₂] (e : M₁ ≃L[R₁] M₂) (c : R₁) (x : M₁) : e (c • x) = c • e x := - (e : M₁ →L[R₁] M₂).map_smul c x - -theorem map_eq_zero_iff (e : M₁ ≃SL[σ₁₂] M₂) {x : M₁} : e x = 0 ↔ x = 0 := - e.toLinearEquiv.map_eq_zero_iff - -attribute [continuity] - ContinuousLinearEquiv.continuous_toFun ContinuousLinearEquiv.continuous_invFun - -@[continuity] -protected theorem continuous (e : M₁ ≃SL[σ₁₂] M₂) : Continuous (e : M₁ → M₂) := - e.continuous_toFun - -protected theorem continuousOn (e : M₁ ≃SL[σ₁₂] M₂) {s : Set M₁} : ContinuousOn (e : M₁ → M₂) s := - e.continuous.continuousOn - -protected theorem continuousAt (e : M₁ ≃SL[σ₁₂] M₂) {x : M₁} : ContinuousAt (e : M₁ → M₂) x := - e.continuous.continuousAt - -protected theorem continuousWithinAt (e : M₁ ≃SL[σ₁₂] M₂) {s : Set M₁} {x : M₁} : - ContinuousWithinAt (e : M₁ → M₂) s x := - e.continuous.continuousWithinAt - -theorem comp_continuousOn_iff {α : Type*} [TopologicalSpace α] (e : M₁ ≃SL[σ₁₂] M₂) {f : α → M₁} - {s : Set α} : ContinuousOn (e ∘ f) s ↔ ContinuousOn f s := - e.toHomeomorph.comp_continuousOn_iff _ _ - -theorem comp_continuous_iff {α : Type*} [TopologicalSpace α] (e : M₁ ≃SL[σ₁₂] M₂) {f : α → M₁} : - Continuous (e ∘ f) ↔ Continuous f := - e.toHomeomorph.comp_continuous_iff - -/-- An extensionality lemma for `R ≃L[R] M`. -/ -theorem ext₁ [TopologicalSpace R₁] {f g : R₁ ≃L[R₁] M₁} (h : f 1 = g 1) : f = g := - ext <| funext fun x => mul_one x ▸ by rw [← smul_eq_mul, map_smul, h, map_smul] - -section - -variable (R₁ M₁) - -/-- The identity map as a continuous linear equivalence. -/ -@[refl] -protected def refl : M₁ ≃L[R₁] M₁ := - { LinearEquiv.refl R₁ M₁ with - continuous_toFun := continuous_id - continuous_invFun := continuous_id } - -@[simp] -theorem refl_apply (x : M₁) : - ContinuousLinearEquiv.refl R₁ M₁ x = x := rfl - -end - -@[simp, norm_cast] -theorem coe_refl : ↑(ContinuousLinearEquiv.refl R₁ M₁) = ContinuousLinearMap.id R₁ M₁ := - rfl - -@[simp, norm_cast] -theorem coe_refl' : ⇑(ContinuousLinearEquiv.refl R₁ M₁) = id := - rfl - -/-- The inverse of a continuous linear equivalence as a continuous linear equivalence -/ -@[symm] -protected def symm (e : M₁ ≃SL[σ₁₂] M₂) : M₂ ≃SL[σ₂₁] M₁ := - { e.toLinearEquiv.symm with - continuous_toFun := e.continuous_invFun - continuous_invFun := e.continuous_toFun } - -@[simp] -theorem symm_toLinearEquiv (e : M₁ ≃SL[σ₁₂] M₂) : e.symm.toLinearEquiv = e.toLinearEquiv.symm := by - ext - rfl - -@[simp] -theorem symm_toHomeomorph (e : M₁ ≃SL[σ₁₂] M₂) : e.toHomeomorph.symm = e.symm.toHomeomorph := - rfl - -/-- See Note [custom simps projection]. We need to specify this projection explicitly in this case, - because it is a composition of multiple projections. -/ -def Simps.apply (h : M₁ ≃SL[σ₁₂] M₂) : M₁ → M₂ := - h - -/-- See Note [custom simps projection] -/ -def Simps.symm_apply (h : M₁ ≃SL[σ₁₂] M₂) : M₂ → M₁ := - h.symm - -initialize_simps_projections ContinuousLinearEquiv (toFun → apply, invFun → symm_apply) - -theorem symm_map_nhds_eq (e : M₁ ≃SL[σ₁₂] M₂) (x : M₁) : map e.symm (𝓝 (e x)) = 𝓝 x := - e.toHomeomorph.symm_map_nhds_eq x - -/-- The composition of two continuous linear equivalences as a continuous linear equivalence. -/ -@[trans] -protected def trans (e₁ : M₁ ≃SL[σ₁₂] M₂) (e₂ : M₂ ≃SL[σ₂₃] M₃) : M₁ ≃SL[σ₁₃] M₃ := - { e₁.toLinearEquiv.trans e₂.toLinearEquiv with - continuous_toFun := e₂.continuous_toFun.comp e₁.continuous_toFun - continuous_invFun := e₁.continuous_invFun.comp e₂.continuous_invFun } - -@[simp] -theorem trans_toLinearEquiv (e₁ : M₁ ≃SL[σ₁₂] M₂) (e₂ : M₂ ≃SL[σ₂₃] M₃) : - (e₁.trans e₂).toLinearEquiv = e₁.toLinearEquiv.trans e₂.toLinearEquiv := by - ext - rfl - -/-- Product of two continuous linear equivalences. The map comes from `Equiv.prodCongr`. -/ -def prod [Module R₁ M₂] [Module R₁ M₃] [Module R₁ M₄] (e : M₁ ≃L[R₁] M₂) (e' : M₃ ≃L[R₁] M₄) : - (M₁ × M₃) ≃L[R₁] M₂ × M₄ := - { e.toLinearEquiv.prod e'.toLinearEquiv with - continuous_toFun := e.continuous_toFun.prodMap e'.continuous_toFun - continuous_invFun := e.continuous_invFun.prodMap e'.continuous_invFun } - -@[simp, norm_cast] -theorem prod_apply [Module R₁ M₂] [Module R₁ M₃] [Module R₁ M₄] (e : M₁ ≃L[R₁] M₂) - (e' : M₃ ≃L[R₁] M₄) (x) : e.prod e' x = (e x.1, e' x.2) := - rfl - -@[simp, norm_cast] -theorem coe_prod [Module R₁ M₂] [Module R₁ M₃] [Module R₁ M₄] (e : M₁ ≃L[R₁] M₂) - (e' : M₃ ≃L[R₁] M₄) : - (e.prod e' : M₁ × M₃ →L[R₁] M₂ × M₄) = (e : M₁ →L[R₁] M₂).prodMap (e' : M₃ →L[R₁] M₄) := - rfl - -theorem prod_symm [Module R₁ M₂] [Module R₁ M₃] [Module R₁ M₄] (e : M₁ ≃L[R₁] M₂) - (e' : M₃ ≃L[R₁] M₄) : (e.prod e').symm = e.symm.prod e'.symm := - rfl - -variable (R₁ M₁ M₂) - -/-- Product of modules is commutative up to continuous linear isomorphism. -/ -@[simps! apply toLinearEquiv] -def prodComm [Module R₁ M₂] : (M₁ × M₂) ≃L[R₁] M₂ × M₁ := - { LinearEquiv.prodComm R₁ M₁ M₂ with - continuous_toFun := continuous_swap - continuous_invFun := continuous_swap } - -@[simp] lemma prodComm_symm [Module R₁ M₂] : (prodComm R₁ M₁ M₂).symm = prodComm R₁ M₂ M₁ := rfl - -variable {R₁ M₁ M₂} - -protected theorem bijective (e : M₁ ≃SL[σ₁₂] M₂) : Function.Bijective e := - e.toLinearEquiv.toEquiv.bijective - -protected theorem injective (e : M₁ ≃SL[σ₁₂] M₂) : Function.Injective e := - e.toLinearEquiv.toEquiv.injective - -protected theorem surjective (e : M₁ ≃SL[σ₁₂] M₂) : Function.Surjective e := - e.toLinearEquiv.toEquiv.surjective - -@[simp] -theorem trans_apply (e₁ : M₁ ≃SL[σ₁₂] M₂) (e₂ : M₂ ≃SL[σ₂₃] M₃) (c : M₁) : - (e₁.trans e₂) c = e₂ (e₁ c) := - rfl - -@[simp] -theorem apply_symm_apply (e : M₁ ≃SL[σ₁₂] M₂) (c : M₂) : e (e.symm c) = c := - e.1.right_inv c - -@[simp] -theorem symm_apply_apply (e : M₁ ≃SL[σ₁₂] M₂) (b : M₁) : e.symm (e b) = b := - e.1.left_inv b - -@[simp] -theorem symm_trans_apply (e₁ : M₂ ≃SL[σ₂₁] M₁) (e₂ : M₃ ≃SL[σ₃₂] M₂) (c : M₁) : - (e₂.trans e₁).symm c = e₂.symm (e₁.symm c) := - rfl - -@[simp] -theorem symm_image_image (e : M₁ ≃SL[σ₁₂] M₂) (s : Set M₁) : e.symm '' (e '' s) = s := - e.toLinearEquiv.toEquiv.symm_image_image s - -@[simp] -theorem image_symm_image (e : M₁ ≃SL[σ₁₂] M₂) (s : Set M₂) : e '' (e.symm '' s) = s := - e.symm.symm_image_image s - -@[simp, norm_cast] -theorem comp_coe (f : M₁ ≃SL[σ₁₂] M₂) (f' : M₂ ≃SL[σ₂₃] M₃) : - (f' : M₂ →SL[σ₂₃] M₃).comp (f : M₁ →SL[σ₁₂] M₂) = (f.trans f' : M₁ →SL[σ₁₃] M₃) := - rfl - --- Porting note: The priority should be higher than `comp_coe`. -@[simp high] -theorem coe_comp_coe_symm (e : M₁ ≃SL[σ₁₂] M₂) : - (e : M₁ →SL[σ₁₂] M₂).comp (e.symm : M₂ →SL[σ₂₁] M₁) = ContinuousLinearMap.id R₂ M₂ := - ContinuousLinearMap.ext e.apply_symm_apply - --- Porting note: The priority should be higher than `comp_coe`. -@[simp high] -theorem coe_symm_comp_coe (e : M₁ ≃SL[σ₁₂] M₂) : - (e.symm : M₂ →SL[σ₂₁] M₁).comp (e : M₁ →SL[σ₁₂] M₂) = ContinuousLinearMap.id R₁ M₁ := - ContinuousLinearMap.ext e.symm_apply_apply - -@[simp] -theorem symm_comp_self (e : M₁ ≃SL[σ₁₂] M₂) : (e.symm : M₂ → M₁) ∘ (e : M₁ → M₂) = id := by - ext x - exact symm_apply_apply e x - -@[simp] -theorem self_comp_symm (e : M₁ ≃SL[σ₁₂] M₂) : (e : M₁ → M₂) ∘ (e.symm : M₂ → M₁) = id := by - ext x - exact apply_symm_apply e x - -@[simp] -theorem symm_symm (e : M₁ ≃SL[σ₁₂] M₂) : e.symm.symm = e := rfl - -@[simp] -theorem refl_symm : (ContinuousLinearEquiv.refl R₁ M₁).symm = ContinuousLinearEquiv.refl R₁ M₁ := - rfl - -theorem symm_symm_apply (e : M₁ ≃SL[σ₁₂] M₂) (x : M₁) : e.symm.symm x = e x := - rfl - -theorem symm_apply_eq (e : M₁ ≃SL[σ₁₂] M₂) {x y} : e.symm x = y ↔ x = e y := - e.toLinearEquiv.symm_apply_eq - -theorem eq_symm_apply (e : M₁ ≃SL[σ₁₂] M₂) {x y} : y = e.symm x ↔ e y = x := - e.toLinearEquiv.eq_symm_apply - -protected theorem image_eq_preimage (e : M₁ ≃SL[σ₁₂] M₂) (s : Set M₁) : e '' s = e.symm ⁻¹' s := - e.toLinearEquiv.toEquiv.image_eq_preimage s - -protected theorem image_symm_eq_preimage (e : M₁ ≃SL[σ₁₂] M₂) (s : Set M₂) : - e.symm '' s = e ⁻¹' s := by rw [e.symm.image_eq_preimage, e.symm_symm] - -@[simp] -protected theorem symm_preimage_preimage (e : M₁ ≃SL[σ₁₂] M₂) (s : Set M₂) : - e.symm ⁻¹' (e ⁻¹' s) = s := - e.toLinearEquiv.toEquiv.symm_preimage_preimage s - -@[simp] -protected theorem preimage_symm_preimage (e : M₁ ≃SL[σ₁₂] M₂) (s : Set M₁) : - e ⁻¹' (e.symm ⁻¹' s) = s := - e.symm.symm_preimage_preimage s - -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₂) : IsUniformEmbedding e := - e.toLinearEquiv.toEquiv.isUniformEmbedding e.toContinuousLinearMap.uniformContinuous - e.symm.toContinuousLinearMap.uniformContinuous - -@[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) : 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. See also `equivOfInverse'`. -/ -def equivOfInverse (f₁ : M₁ →SL[σ₁₂] M₂) (f₂ : M₂ →SL[σ₂₁] M₁) (h₁ : Function.LeftInverse f₂ f₁) - (h₂ : Function.RightInverse f₂ f₁) : M₁ ≃SL[σ₁₂] M₂ := - { f₁ with - continuous_toFun := f₁.continuous - invFun := f₂ - continuous_invFun := f₂.continuous - left_inv := h₁ - right_inv := h₂ } - -@[simp] -theorem equivOfInverse_apply (f₁ : M₁ →SL[σ₁₂] M₂) (f₂ h₁ h₂ x) : - equivOfInverse f₁ f₂ h₁ h₂ x = f₁ x := - rfl - -@[simp] -theorem symm_equivOfInverse (f₁ : M₁ →SL[σ₁₂] M₂) (f₂ h₁ h₂) : - (equivOfInverse f₁ f₂ h₁ h₂).symm = equivOfInverse f₂ f₁ h₂ h₁ := - rfl - -/-- Create a `ContinuousLinearEquiv` from two `ContinuousLinearMap`s that are -inverse of each other, in the `ContinuousLinearMap.comp` sense. See also `equivOfInverse`. -/ -def equivOfInverse' (f₁ : M₁ →SL[σ₁₂] M₂) (f₂ : M₂ →SL[σ₂₁] M₁) - (h₁ : f₁.comp f₂ = .id R₂ M₂) (h₂ : f₂.comp f₁ = .id R₁ M₁) : M₁ ≃SL[σ₁₂] M₂ := - equivOfInverse f₁ f₂ - (fun x ↦ by simpa using congr($(h₂) x)) (fun x ↦ by simpa using congr($(h₁) x)) - -@[simp] -theorem equivOfInverse'_apply (f₁ : M₁ →SL[σ₁₂] M₂) (f₂ h₁ h₂ x) : - equivOfInverse' f₁ f₂ h₁ h₂ x = f₁ x := - rfl - -/-- The inverse of `equivOfInverse'` is obtained by swapping the order of its parameters. -/ -@[simp] -theorem symm_equivOfInverse' (f₁ : M₁ →SL[σ₁₂] M₂) (f₂ h₁ h₂) : - (equivOfInverse' f₁ f₂ h₁ h₂).symm = equivOfInverse' f₂ f₁ h₂ h₁ := - rfl - -variable (M₁) - -/-- The continuous linear equivalences from `M` to itself form a group under composition. -/ -instance automorphismGroup : Group (M₁ ≃L[R₁] M₁) where - mul f g := g.trans f - one := ContinuousLinearEquiv.refl R₁ M₁ - inv f := f.symm - mul_assoc f g h := by - ext - rfl - mul_one f := by - ext - rfl - one_mul f := by - ext - rfl - inv_mul_cancel f := by - ext x - exact f.left_inv x - -variable {M₁} {R₄ : Type*} [Semiring R₄] [Module R₄ M₄] {σ₃₄ : R₃ →+* R₄} {σ₄₃ : R₄ →+* R₃} - [RingHomInvPair σ₃₄ σ₄₃] [RingHomInvPair σ₄₃ σ₃₄] {σ₂₄ : R₂ →+* R₄} {σ₁₄ : R₁ →+* R₄} - [RingHomCompTriple σ₂₁ σ₁₄ σ₂₄] [RingHomCompTriple σ₂₄ σ₄₃ σ₂₃] [RingHomCompTriple σ₁₃ σ₃₄ σ₁₄] - -/-- The continuous linear equivalence between `ULift M₁` and `M₁`. - -This is a continuous version of `ULift.moduleEquiv`. -/ -def ulift : ULift M₁ ≃L[R₁] M₁ := - { ULift.moduleEquiv with - continuous_toFun := continuous_uLift_down - continuous_invFun := continuous_uLift_up } - -/-- A pair of continuous (semi)linear equivalences generates an equivalence between the spaces of -continuous linear maps. See also `ContinuousLinearEquiv.arrowCongr`. -/ -@[simps] -def arrowCongrEquiv (e₁₂ : M₁ ≃SL[σ₁₂] M₂) (e₄₃ : M₄ ≃SL[σ₄₃] M₃) : - (M₁ →SL[σ₁₄] M₄) ≃ (M₂ →SL[σ₂₃] M₃) where - toFun f := (e₄₃ : M₄ →SL[σ₄₃] M₃).comp (f.comp (e₁₂.symm : M₂ →SL[σ₂₁] M₁)) - invFun f := (e₄₃.symm : M₃ →SL[σ₃₄] M₄).comp (f.comp (e₁₂ : M₁ →SL[σ₁₂] M₂)) - left_inv f := - ContinuousLinearMap.ext fun x => by - simp only [ContinuousLinearMap.comp_apply, symm_apply_apply, coe_coe] - right_inv f := - ContinuousLinearMap.ext fun x => by - simp only [ContinuousLinearMap.comp_apply, apply_symm_apply, coe_coe] - -section piCongrRight - -variable {ι : Type*} {M : ι → Type*} [∀ i, TopologicalSpace (M i)] [∀ i, AddCommMonoid (M i)] - [∀ i, Module R₁ (M i)] {N : ι → Type*} [∀ i, TopologicalSpace (N i)] [∀ i, AddCommMonoid (N i)] - [∀ i, Module R₁ (N i)] (f : (i : ι) → M i ≃L[R₁] N i) - -/-- Combine a family of continuous linear equivalences into a continuous linear equivalence of -pi-types. -/ -def piCongrRight : ((i : ι) → M i) ≃L[R₁] (i : ι) → N i := - { LinearEquiv.piCongrRight fun i ↦ f i with - continuous_toFun := by - exact continuous_pi fun i ↦ (f i).continuous_toFun.comp (continuous_apply i) - continuous_invFun := by - exact continuous_pi fun i => (f i).continuous_invFun.comp (continuous_apply i) } - -@[simp] -theorem piCongrRight_apply (m : (i : ι) → M i) (i : ι) : - piCongrRight f m i = (f i) (m i) := rfl - -@[simp] -theorem piCongrRight_symm_apply (n : (i : ι) → N i) (i : ι) : - (piCongrRight f).symm n i = (f i).symm (n i) := rfl - -end piCongrRight - -end AddCommMonoid - -section AddCommGroup - -variable {R : Type*} [Semiring R] {M : Type*} [TopologicalSpace M] [AddCommGroup M] {M₂ : Type*} - [TopologicalSpace M₂] [AddCommGroup M₂] {M₃ : Type*} [TopologicalSpace M₃] [AddCommGroup M₃] - {M₄ : Type*} [TopologicalSpace M₄] [AddCommGroup M₄] [Module R M] [Module R M₂] [Module R M₃] - [Module R M₄] - -variable [TopologicalAddGroup M₄] - -/-- Equivalence given by a block lower diagonal matrix. `e` and `e'` are diagonal square blocks, - and `f` is a rectangular block below the diagonal. -/ -def skewProd (e : M ≃L[R] M₂) (e' : M₃ ≃L[R] M₄) (f : M →L[R] M₄) : (M × M₃) ≃L[R] M₂ × M₄ := - { - e.toLinearEquiv.skewProd e'.toLinearEquiv - ↑f with - continuous_toFun := - (e.continuous_toFun.comp continuous_fst).prod_mk - ((e'.continuous_toFun.comp continuous_snd).add <| f.continuous.comp continuous_fst) - continuous_invFun := - (e.continuous_invFun.comp continuous_fst).prod_mk - (e'.continuous_invFun.comp <| - continuous_snd.sub <| f.continuous.comp <| e.continuous_invFun.comp continuous_fst) } - -@[simp] -theorem skewProd_apply (e : M ≃L[R] M₂) (e' : M₃ ≃L[R] M₄) (f : M →L[R] M₄) (x) : - e.skewProd e' f x = (e x.1, e' x.2 + f x.1) := - rfl - -@[simp] -theorem skewProd_symm_apply (e : M ≃L[R] M₂) (e' : M₃ ≃L[R] M₄) (f : M →L[R] M₄) (x) : - (e.skewProd e' f).symm x = (e.symm x.1, e'.symm (x.2 - f (e.symm x.1))) := - rfl - -variable (R) in -/-- The negation map as a continuous linear equivalence. -/ -def neg [ContinuousNeg M] : - M ≃L[R] M := - { LinearEquiv.neg R with - continuous_toFun := continuous_neg - continuous_invFun := continuous_neg } - -@[simp] -theorem coe_neg [ContinuousNeg M] : - (neg R : M → M) = -id := rfl - -@[simp] -theorem neg_apply [ContinuousNeg M] (x : M) : - neg R x = -x := by simp - -@[simp] -theorem symm_neg [ContinuousNeg M] : - (neg R : M ≃L[R] M).symm = neg R := rfl - -end AddCommGroup - -section Ring - -variable {R : Type*} [Ring R] {R₂ : Type*} [Ring R₂] {M : Type*} [TopologicalSpace M] - [AddCommGroup M] [Module R M] {M₂ : Type*} [TopologicalSpace M₂] [AddCommGroup M₂] [Module R₂ M₂] - -variable {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] - -theorem map_sub (e : M ≃SL[σ₁₂] M₂) (x y : M) : e (x - y) = e x - e y := - (e : M →SL[σ₁₂] M₂).map_sub x y - -theorem map_neg (e : M ≃SL[σ₁₂] M₂) (x : M) : e (-x) = -e x := - (e : M →SL[σ₁₂] M₂).map_neg x - -section - -/-! The next theorems cover the identification between `M ≃L[R] M`and the group of units of the ring -`M →L[R] M`. -/ - -/-- An invertible continuous linear map `f` determines a continuous equivalence from `M` to itself. --/ -def ofUnit (f : (M →L[R] M)ˣ) : M ≃L[R] M where - toLinearEquiv := - { toFun := f.val - map_add' := by simp - map_smul' := by simp - invFun := f.inv - left_inv := fun x => - show (f.inv * f.val) x = x by - rw [f.inv_val] - simp - right_inv := fun x => - show (f.val * f.inv) x = x by - rw [f.val_inv] - simp } - continuous_toFun := f.val.continuous - continuous_invFun := f.inv.continuous - -/-- A continuous equivalence from `M` to itself determines an invertible continuous linear map. -/ -def toUnit (f : M ≃L[R] M) : (M →L[R] M)ˣ where - val := f - inv := f.symm - val_inv := by - ext - simp - inv_val := by - ext - simp - -variable (R M) - -/-- The units of the algebra of continuous `R`-linear endomorphisms of `M` is multiplicatively -equivalent to the type of continuous linear equivalences between `M` and itself. -/ -def unitsEquiv : (M →L[R] M)ˣ ≃* M ≃L[R] M where - toFun := ofUnit - invFun := toUnit - left_inv f := by - ext - rfl - right_inv f := by - ext - rfl - map_mul' x y := by - ext - rfl - -@[simp] -theorem unitsEquiv_apply (f : (M →L[R] M)ˣ) (x : M) : unitsEquiv R M f x = (f : M →L[R] M) x := - rfl - -end - -section - -variable (R) [TopologicalSpace R] -variable [ContinuousMul R] - -/-- Continuous linear equivalences `R ≃L[R] R` are enumerated by `Rˣ`. -/ -def unitsEquivAut : Rˣ ≃ R ≃L[R] R where - toFun u := - equivOfInverse (ContinuousLinearMap.smulRight (1 : R →L[R] R) ↑u) - (ContinuousLinearMap.smulRight (1 : R →L[R] R) ↑u⁻¹) (fun x => by simp) fun x => by simp - invFun e := - ⟨e 1, e.symm 1, by rw [← smul_eq_mul, ← map_smul, smul_eq_mul, mul_one, symm_apply_apply], by - rw [← smul_eq_mul, ← map_smul, smul_eq_mul, mul_one, apply_symm_apply]⟩ - left_inv u := Units.ext <| by simp - right_inv e := ext₁ <| by simp - -variable {R} - -@[simp] -theorem unitsEquivAut_apply (u : Rˣ) (x : R) : unitsEquivAut R u x = x * u := - rfl - -@[simp] -theorem unitsEquivAut_apply_symm (u : Rˣ) (x : R) : (unitsEquivAut R u).symm x = x * ↑u⁻¹ := - rfl - -@[simp] -theorem unitsEquivAut_symm_apply (e : R ≃L[R] R) : ↑((unitsEquivAut R).symm e) = e 1 := - rfl - -end - -variable [Module R M₂] [TopologicalAddGroup M] - -/-- A pair of continuous linear maps such that `f₁ ∘ f₂ = id` generates a continuous -linear equivalence `e` between `M` and `M₂ × f₁.ker` such that `(e x).2 = x` for `x ∈ f₁.ker`, -`(e x).1 = f₁ x`, and `(e (f₂ y)).2 = 0`. The map is given by `e x = (f₁ x, x - f₂ (f₁ x))`. -/ -def equivOfRightInverse (f₁ : M →L[R] M₂) (f₂ : M₂ →L[R] M) (h : Function.RightInverse f₂ f₁) : - M ≃L[R] M₂ × ker f₁ := - equivOfInverse (f₁.prod (f₁.projKerOfRightInverse f₂ h)) (f₂.coprod (ker f₁).subtypeL) - (fun x => by simp) fun ⟨x, y⟩ => by - -- Porting note: `simp` timeouts. - rw [ContinuousLinearMap.coprod_apply, - Submodule.subtypeL_apply, _root_.map_add, ContinuousLinearMap.prod_apply, h x, - ContinuousLinearMap.projKerOfRightInverse_comp_inv, - ContinuousLinearMap.prod_apply, LinearMap.map_coe_ker, - ContinuousLinearMap.projKerOfRightInverse_apply_idem, Prod.mk_add_mk, add_zero, zero_add] - -@[simp] -theorem fst_equivOfRightInverse (f₁ : M →L[R] M₂) (f₂ : M₂ →L[R] M) - (h : Function.RightInverse f₂ f₁) (x : M) : (equivOfRightInverse f₁ f₂ h x).1 = f₁ x := - rfl - -@[simp] -theorem snd_equivOfRightInverse (f₁ : M →L[R] M₂) (f₂ : M₂ →L[R] M) - (h : Function.RightInverse f₂ f₁) (x : M) : - ((equivOfRightInverse f₁ f₂ h x).2 : M) = x - f₂ (f₁ x) := - rfl - -@[simp] -theorem equivOfRightInverse_symm_apply (f₁ : M →L[R] M₂) (f₂ : M₂ →L[R] M) - (h : Function.RightInverse f₂ f₁) (y : M₂ × ker f₁) : - (equivOfRightInverse f₁ f₂ h).symm y = f₂ y.1 + y.2 := - rfl - -end Ring - -section - -variable (ι R M : Type*) [Unique ι] [Semiring R] [AddCommMonoid M] [Module R M] - [TopologicalSpace M] - -/-- If `ι` has a unique element, then `ι → M` is continuously linear equivalent to `M`. -/ -def funUnique : (ι → M) ≃L[R] M := - { Homeomorph.funUnique ι M with toLinearEquiv := LinearEquiv.funUnique ι R M } - -variable {ι R M} - -@[simp] -theorem coe_funUnique : ⇑(funUnique ι R M) = Function.eval default := - rfl - -@[simp] -theorem coe_funUnique_symm : ⇑(funUnique ι R M).symm = Function.const ι := - rfl - -variable (R M) - -/-- Continuous linear equivalence between dependent functions `(i : Fin 2) → M i` and `M 0 × M 1`. --/ -@[simps! (config := .asFn) apply symm_apply] -def piFinTwo (M : Fin 2 → Type*) [∀ i, AddCommMonoid (M i)] [∀ i, Module R (M i)] - [∀ i, TopologicalSpace (M i)] : ((i : _) → M i) ≃L[R] M 0 × M 1 := - { Homeomorph.piFinTwo M with toLinearEquiv := LinearEquiv.piFinTwo R M } - -/-- Continuous linear equivalence between vectors in `M² = Fin 2 → M` and `M × M`. -/ -@[simps! (config := .asFn) apply symm_apply] -def finTwoArrow : (Fin 2 → M) ≃L[R] M × M := - { piFinTwo R fun _ => M with toLinearEquiv := LinearEquiv.finTwoArrow R M } - -end - -end ContinuousLinearEquiv - -namespace ContinuousLinearMap - -variable {R : Type*} {M M₂ M₃ : Type*} - [TopologicalSpace M] [TopologicalSpace M₂] [TopologicalSpace M₃] - -section - -variable [Semiring R] - [AddCommMonoid M] [Module R M] - [AddCommMonoid M₂] [Module R M₂] - [AddCommMonoid M₃] [Module R M₃] - -/-- A continuous linear map is invertible if it is the forward direction of a continuous linear -equivalence. -/ -def IsInvertible (f : M →L[R] M₂) : Prop := - ∃ (A : M ≃L[R] M₂), A = f - -open Classical in -/-- Introduce a function `inverse` from `M →L[R] M₂` to `M₂ →L[R] M`, which sends `f` to `f.symm` if -`f` is a continuous linear equivalence and to `0` otherwise. This definition is somewhat ad hoc, -but one needs a fully (rather than partially) defined inverse function for some purposes, including -for calculus. -/ -noncomputable def inverse : (M →L[R] M₂) → M₂ →L[R] M := fun f => - if h : f.IsInvertible then ((Classical.choose h).symm : M₂ →L[R] M) else 0 - -@[simp] lemma isInvertible_equiv {f : M ≃L[R] M₂} : IsInvertible (f : M →L[R] M₂) := ⟨f, rfl⟩ - -/-- By definition, if `f` is invertible then `inverse f = f.symm`. -/ -@[simp] -theorem inverse_equiv (e : M ≃L[R] M₂) : inverse (e : M →L[R] M₂) = e.symm := by - simp [inverse] - -/-- By definition, if `f` is not invertible then `inverse f = 0`. -/ -@[simp] lemma inverse_of_not_isInvertible - {f : M →L[R] M₂} (hf : ¬ f.IsInvertible) : f.inverse = 0 := - dif_neg hf - -@[deprecated (since := "2024-10-29")] alias inverse_non_equiv := inverse_of_not_isInvertible - -@[simp] theorem inverse_zero : inverse (0 : M →L[R] M₂) = 0 := by - by_cases h : IsInvertible (0 : M →L[R] M₂) - · rcases h with ⟨e', he'⟩ - simp only [← he', inverse_equiv] - ext v - apply e'.injective - rw [← ContinuousLinearEquiv.coe_coe, he'] - rfl - · exact inverse_of_not_isInvertible h - -lemma IsInvertible.comp {g : M₂ →L[R] M₃} {f : M →L[R] M₂} - (hg : g.IsInvertible) (hf : f.IsInvertible) : (g ∘L f).IsInvertible := by - rcases hg with ⟨N, rfl⟩ - rcases hf with ⟨M, rfl⟩ - exact ⟨M.trans N, rfl⟩ - -lemma IsInvertible.of_inverse {f : M →L[R] M₂} {g : M₂ →L[R] M} - (hf : f ∘L g = id R M₂) (hg : g ∘L f = id R M) : - f.IsInvertible := - ⟨ContinuousLinearEquiv.equivOfInverse' _ _ hf hg, rfl⟩ - -lemma inverse_eq {f : M →L[R] M₂} {g : M₂ →L[R] M} (hf : f ∘L g = id R M₂) (hg : g ∘L f = id R M) : - f.inverse = g := by - have : f = ContinuousLinearEquiv.equivOfInverse' f g hf hg := rfl - rw [this, inverse_equiv] - rfl - -lemma IsInvertible.inverse_apply_eq {f : M →L[R] M₂} {x : M} {y : M₂} (hf : f.IsInvertible) : - f.inverse y = x ↔ y = f x := by - rcases hf with ⟨M, rfl⟩ - simp only [inverse_equiv, ContinuousLinearEquiv.coe_coe] - exact ContinuousLinearEquiv.symm_apply_eq M - -@[simp] lemma isInvertible_equiv_comp {e : M₂ ≃L[R] M₃} {f : M →L[R] M₂} : - ((e : M₂ →L[R] M₃) ∘L f).IsInvertible ↔ f.IsInvertible := by - constructor - · rintro ⟨A, hA⟩ - have : f = e.symm ∘L ((e : M₂ →L[R] M₃) ∘L f) := by ext; simp - rw [this, ← hA] - simp - · rintro ⟨M, rfl⟩ - simp - -@[simp] lemma isInvertible_comp_equiv {e : M₃ ≃L[R] M} {f : M →L[R] M₂} : - (f ∘L (e : M₃ →L[R] M)).IsInvertible ↔ f.IsInvertible := by - constructor - · rintro ⟨A, hA⟩ - have : f = (f ∘L (e : M₃ →L[R] M)) ∘L e.symm := by ext; simp - rw [this, ← hA] - simp - · rintro ⟨M, rfl⟩ - simp - -@[simp] lemma inverse_equiv_comp {e : M₂ ≃L[R] M₃} {f : M →L[R] M₂} : - (e ∘L f).inverse = f.inverse ∘L (e.symm : M₃ →L[R] M₂) := by - by_cases hf : f.IsInvertible - · rcases hf with ⟨A, rfl⟩ - simp only [ContinuousLinearEquiv.comp_coe, inverse_equiv, ContinuousLinearEquiv.coe_inj] - rfl - · rw [inverse_of_not_isInvertible (by simp [hf]), inverse_of_not_isInvertible hf] - rfl - -@[simp] lemma inverse_comp_equiv {e : M₃ ≃L[R] M} {f : M →L[R] M₂} : - (f ∘L e).inverse = (e.symm : M →L[R] M₃) ∘L f.inverse := by - by_cases hf : f.IsInvertible - · rcases hf with ⟨A, rfl⟩ - simp only [ContinuousLinearEquiv.comp_coe, inverse_equiv, ContinuousLinearEquiv.coe_inj] - rfl - · rw [inverse_of_not_isInvertible (by simp [hf]), inverse_of_not_isInvertible hf] - simp - -lemma IsInvertible.inverse_comp_of_left {g : M₂ →L[R] M₃} {f : M →L[R] M₂} - (hg : g.IsInvertible) : (g ∘L f).inverse = f.inverse ∘L g.inverse := by - rcases hg with ⟨N, rfl⟩ - simp - -lemma IsInvertible.inverse_comp_apply_of_left {g : M₂ →L[R] M₃} {f : M →L[R] M₂} {v : M₃} - (hg : g.IsInvertible) : (g ∘L f).inverse v = f.inverse (g.inverse v) := by - simp only [hg.inverse_comp_of_left, coe_comp', Function.comp_apply] - -lemma IsInvertible.inverse_comp_of_right {g : M₂ →L[R] M₃} {f : M →L[R] M₂} - (hf : f.IsInvertible) : (g ∘L f).inverse = f.inverse ∘L g.inverse := by - rcases hf with ⟨M, rfl⟩ - simp - -lemma IsInvertible.inverse_comp_apply_of_right {g : M₂ →L[R] M₃} {f : M →L[R] M₂} {v : M₃} - (hf : f.IsInvertible) : (g ∘L f).inverse v = f.inverse (g.inverse v) := by - simp only [hf.inverse_comp_of_right, coe_comp', Function.comp_apply] - -end - -section - -variable [Ring R] -variable [AddCommGroup M] [Module R M] -variable [AddCommGroup M₂] [Module R M₂] - -@[simp] -theorem ring_inverse_equiv (e : M ≃L[R] M) : Ring.inverse ↑e = inverse (e : M →L[R] M) := by - suffices Ring.inverse ((ContinuousLinearEquiv.unitsEquiv _ _).symm e : M →L[R] M) = inverse ↑e by - convert this - simp - rfl - -/-- The function `ContinuousLinearEquiv.inverse` can be written in terms of `Ring.inverse` for the -ring of self-maps of the domain. -/ -theorem to_ring_inverse (e : M ≃L[R] M₂) (f : M →L[R] M₂) : - inverse f = Ring.inverse ((e.symm : M₂ →L[R] M).comp f) ∘L e.symm := by - by_cases h₁ : f.IsInvertible - · obtain ⟨e', he'⟩ := h₁ - rw [← he'] - change _ = Ring.inverse (e'.trans e.symm : M →L[R] M) ∘L (e.symm : M₂ →L[R] M) - ext - simp - · suffices ¬IsUnit ((e.symm : M₂ →L[R] M).comp f) by simp [this, h₁] - contrapose! h₁ - rcases h₁ with ⟨F, hF⟩ - use (ContinuousLinearEquiv.unitsEquiv _ _ F).trans e - ext - dsimp - rw [hF] - simp - -theorem ring_inverse_eq_map_inverse : Ring.inverse = @inverse R M M _ _ _ _ _ _ _ := by - ext - simp [to_ring_inverse (ContinuousLinearEquiv.refl R M)] - -@[simp] theorem inverse_id : (id R M).inverse = id R M := by - rw [← ring_inverse_eq_map_inverse] - exact Ring.inverse_one _ - -end - -end ContinuousLinearMap - -namespace Submodule - -variable {R : Type*} [Ring R] {M : Type*} [TopologicalSpace M] [AddCommGroup M] [Module R M] - -open ContinuousLinearMap - -/-- A submodule `p` is called *complemented* if there exists a continuous projection `M →ₗ[R] p`. -/ -def ClosedComplemented (p : Submodule R M) : Prop := - ∃ f : M →L[R] p, ∀ x : p, f x = x - -theorem ClosedComplemented.exists_isClosed_isCompl {p : Submodule R M} [T1Space p] - (h : ClosedComplemented p) : - ∃ q : Submodule R M, IsClosed (q : Set M) ∧ IsCompl p q := - Exists.elim h fun f hf => ⟨ker f, isClosed_ker f, LinearMap.isCompl_of_proj hf⟩ - -protected theorem ClosedComplemented.isClosed [TopologicalAddGroup M] [T1Space M] - {p : Submodule R M} (h : ClosedComplemented p) : IsClosed (p : Set M) := by - rcases h with ⟨f, hf⟩ - have : ker (id R M - p.subtypeL.comp f) = p := LinearMap.ker_id_sub_eq_of_proj hf - exact this ▸ isClosed_ker _ - -@[simp] -theorem closedComplemented_bot : ClosedComplemented (⊥ : Submodule R M) := - ⟨0, fun x => by simp only [zero_apply, eq_zero_of_bot_submodule x]⟩ - -@[simp] -theorem closedComplemented_top : ClosedComplemented (⊤ : Submodule R M) := - ⟨(id R M).codRestrict ⊤ fun _x => trivial, fun x => Subtype.ext_iff_val.2 <| by simp⟩ - -/-- If `p` is a closed complemented submodule, -then there exists a submodule `q` and a continuous linear equivalence `M ≃L[R] (p × q)` such that -`e (x : p) = (x, 0)`, `e (y : q) = (0, y)`, and `e.symm x = x.1 + x.2`. - -In fact, the properties of `e` imply the properties of `e.symm` and vice versa, -but we provide both for convenience. -/ -lemma ClosedComplemented.exists_submodule_equiv_prod [TopologicalAddGroup M] - {p : Submodule R M} (hp : p.ClosedComplemented) : - ∃ (q : Submodule R M) (e : M ≃L[R] (p × q)), - (∀ x : p, e x = (x, 0)) ∧ (∀ y : q, e y = (0, y)) ∧ (∀ x, e.symm x = x.1 + x.2) := - let ⟨f, hf⟩ := hp - ⟨LinearMap.ker f, .equivOfRightInverse _ p.subtypeL hf, - fun _ ↦ by ext <;> simp [hf], fun _ ↦ by ext <;> simp [hf], fun _ ↦ rfl⟩ - -end Submodule - -theorem ContinuousLinearMap.closedComplemented_ker_of_rightInverse {R : Type*} [Ring R] - {M : Type*} [TopologicalSpace M] [AddCommGroup M] {M₂ : Type*} [TopologicalSpace M₂] - [AddCommGroup M₂] [Module R M] [Module R M₂] [TopologicalAddGroup M] (f₁ : M →L[R] M₂) - (f₂ : M₂ →L[R] M) (h : Function.RightInverse f₂ f₁) : (ker f₁).ClosedComplemented := - ⟨f₁.projKerOfRightInverse f₂ h, f₁.projKerOfRightInverse_apply_idem f₂ h⟩ - -section Quotient - -namespace Submodule - -variable {R M : Type*} [Ring R] [AddCommGroup M] [Module R M] [TopologicalSpace M] - (S : Submodule R M) - --- Porting note: This is required in Lean4. -instance _root_.QuotientModule.Quotient.topologicalSpace : TopologicalSpace (M ⧸ S) := - inferInstanceAs (TopologicalSpace (Quotient S.quotientRel)) - -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) 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) := - letI : IsClosed (S.toAddSubgroup : Set M) := ‹_› - QuotientAddGroup.instT3Space S.toAddSubgroup - -end Submodule +end Submodule end Quotient - -namespace MulOpposite - -variable (R : Type*) [Semiring R] [τR : TopologicalSpace R] [TopologicalSemiring R] - {M : Type*} [AddCommMonoid M] [Module R M] [TopologicalSpace M] [ContinuousSMul R M] - -/-- The function `op` is a continuous linear equivalence. -/ -@[simps!] -def opContinuousLinearEquiv : M ≃L[R] Mᵐᵒᵖ where - __ := MulOpposite.opLinearEquiv R - -end MulOpposite - -set_option linter.style.longFile 2700 diff --git a/Mathlib/Topology/Algebra/Module/Determinant.lean b/Mathlib/Topology/Algebra/Module/Determinant.lean index 6ec5fa3a057a1..cf06e39304938 100644 --- a/Mathlib/Topology/Algebra/Module/Determinant.lean +++ b/Mathlib/Topology/Algebra/Module/Determinant.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jan-David Salchow, Sébastien Gouëzel, Jean Lo, Yury Kudryashov, Frédéric Dupuis, Heather Macbeth -/ -import Mathlib.Topology.Algebra.Module.Basic +import Mathlib.Topology.Algebra.Module.Equiv import Mathlib.LinearAlgebra.Determinant /-! diff --git a/Mathlib/Topology/Algebra/Module/Equiv.lean b/Mathlib/Topology/Algebra/Module/Equiv.lean new file mode 100644 index 0000000000000..b2d1f0f7870e1 --- /dev/null +++ b/Mathlib/Topology/Algebra/Module/Equiv.lean @@ -0,0 +1,1035 @@ +/- +Copyright (c) 2019 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jan-David Salchow, Sébastien Gouëzel, Jean Lo, Yury Kudryashov, Frédéric Dupuis, + Heather Macbeth +-/ +import Mathlib.Topology.Algebra.Module.LinearMapPiProd + +/-! +# Continuous linear equivalences + +Continuous semilinear / linear / star-linear equivalences between topological modules are denoted +by `M ≃SL[σ] M₂`, `M ≃L[R] M₂` and `M ≃L⋆[R] M₂`. +-/ + +assert_not_exists Star.star + +open LinearMap (ker range) +open Topology Filter Pointwise + +universe u v w u' + +section + +variable {R : Type*} {M : Type*} [Ring R] [TopologicalSpace R] [TopologicalSpace M] + [AddCommGroup M] [ContinuousAdd M] [Module R M] [ContinuousSMul R M] + +variable (R M) + +/-- Continuous linear equivalences between modules. We only put the type classes that are necessary +for the definition, although in applications `M` and `M₂` will be topological modules over the +topological semiring `R`. -/ +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not ported yet; was @[nolint has_nonempty_instance] +structure ContinuousLinearEquiv {R : Type*} {S : Type*} [Semiring R] [Semiring S] (σ : R →+* S) + {σ' : S →+* R} [RingHomInvPair σ σ'] [RingHomInvPair σ' σ] (M : Type*) [TopologicalSpace M] + [AddCommMonoid M] (M₂ : Type*) [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R M] + [Module S M₂] extends M ≃ₛₗ[σ] M₂ where + continuous_toFun : Continuous toFun := by continuity + continuous_invFun : Continuous invFun := by continuity + +attribute [inherit_doc ContinuousLinearEquiv] ContinuousLinearEquiv.continuous_toFun +ContinuousLinearEquiv.continuous_invFun + +@[inherit_doc] +notation:50 M " ≃SL[" σ "] " M₂ => ContinuousLinearEquiv σ M M₂ + +@[inherit_doc] +notation:50 M " ≃L[" R "] " M₂ => ContinuousLinearEquiv (RingHom.id R) M M₂ + +/-- `ContinuousSemilinearEquivClass F σ M M₂` asserts `F` is a type of bundled continuous +`σ`-semilinear equivs `M → M₂`. See also `ContinuousLinearEquivClass F R M M₂` for the case +where `σ` is the identity map on `R`. A map `f` between an `R`-module and an `S`-module over a ring +homomorphism `σ : R →+* S` is semilinear if it satisfies the two properties `f (x + y) = f x + f y` +and `f (c • x) = (σ c) • f x`. -/ +class ContinuousSemilinearEquivClass (F : Type*) {R : outParam Type*} {S : outParam Type*} + [Semiring R] [Semiring S] (σ : outParam <| R →+* S) {σ' : outParam <| S →+* R} + [RingHomInvPair σ σ'] [RingHomInvPair σ' σ] (M : outParam Type*) [TopologicalSpace M] + [AddCommMonoid M] (M₂ : outParam Type*) [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R M] + [Module S M₂] [EquivLike F M M₂] extends SemilinearEquivClass F σ M M₂ : Prop where + map_continuous : ∀ f : F, Continuous f := by continuity + inv_continuous : ∀ f : F, Continuous (EquivLike.inv f) := by continuity + +attribute [inherit_doc ContinuousSemilinearEquivClass] +ContinuousSemilinearEquivClass.map_continuous +ContinuousSemilinearEquivClass.inv_continuous + +/-- `ContinuousLinearEquivClass F σ M M₂` asserts `F` is a type of bundled continuous +`R`-linear equivs `M → M₂`. This is an abbreviation for +`ContinuousSemilinearEquivClass F (RingHom.id R) M M₂`. -/ +abbrev ContinuousLinearEquivClass (F : Type*) (R : outParam Type*) [Semiring R] + (M : outParam Type*) [TopologicalSpace M] [AddCommMonoid M] (M₂ : outParam Type*) + [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R M] [Module R M₂] [EquivLike F M M₂] := + ContinuousSemilinearEquivClass F (RingHom.id R) M M₂ + +namespace ContinuousSemilinearEquivClass + +variable (F : Type*) {R : Type*} {S : Type*} [Semiring R] [Semiring S] (σ : R →+* S) + {σ' : S →+* R} [RingHomInvPair σ σ'] [RingHomInvPair σ' σ] + (M : Type*) [TopologicalSpace M] [AddCommMonoid M] + (M₂ : Type*) [TopologicalSpace M₂] [AddCommMonoid M₂] + [Module R M] [Module S M₂] + +-- `σ'` becomes a metavariable, but it's OK since it's an outparam +instance (priority := 100) continuousSemilinearMapClass [EquivLike F M M₂] + [s : ContinuousSemilinearEquivClass F σ M M₂] : ContinuousSemilinearMapClass F σ M M₂ := + { s with } + +instance (priority := 100) [EquivLike F M M₂] + [s : ContinuousSemilinearEquivClass F σ M M₂] : HomeomorphClass F M M₂ := + { s with } + +end ContinuousSemilinearEquivClass + +namespace ContinuousLinearMap + +section Pi + +variable {R : Type*} [Semiring R] {M : Type*} [TopologicalSpace M] [AddCommMonoid M] [Module R M] + {M₂ : Type*} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R M₂] {ι : Type*} {φ : ι → Type*} + [∀ i, TopologicalSpace (φ i)] [∀ i, AddCommMonoid (φ i)] [∀ i, Module R (φ i)] + +variable (R φ) + +/-- If `I` and `J` are complementary index sets, the product of the kernels of the `J`th projections +of `φ` is linearly equivalent to the product over `I`. -/ +def iInfKerProjEquiv {I J : Set ι} [DecidablePred fun i => i ∈ I] (hd : Disjoint I J) + (hu : Set.univ ⊆ I ∪ J) : + (⨅ i ∈ J, ker (proj i : (∀ i, φ i) →L[R] φ i) : + Submodule R (∀ i, φ i)) ≃L[R] ∀ i : I, φ i where + toLinearEquiv := LinearMap.iInfKerProjEquiv R φ hd hu + continuous_toFun := + continuous_pi fun i => + Continuous.comp (continuous_apply (π := φ) i) <| + @continuous_subtype_val _ _ fun x => + x ∈ (⨅ i ∈ J, ker (proj i : (∀ i, φ i) →L[R] φ i) : Submodule R (∀ i, φ i)) + continuous_invFun := + Continuous.subtype_mk + (continuous_pi fun i => by + dsimp + split_ifs <;> [apply continuous_apply; exact continuous_zero]) + _ + +end Pi + +end ContinuousLinearMap + +namespace ContinuousLinearEquiv + +section AddCommMonoid + +variable {R₁ : Type*} {R₂ : Type*} {R₃ : Type*} [Semiring R₁] [Semiring R₂] [Semiring R₃] + {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] + {σ₂₃ : R₂ →+* R₃} {σ₃₂ : R₃ →+* R₂} [RingHomInvPair σ₂₃ σ₃₂] [RingHomInvPair σ₃₂ σ₂₃] + {σ₁₃ : R₁ →+* R₃} {σ₃₁ : R₃ →+* R₁} [RingHomInvPair σ₁₃ σ₃₁] [RingHomInvPair σ₃₁ σ₁₃] + [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] [RingHomCompTriple σ₃₂ σ₂₁ σ₃₁] {M₁ : Type*} + [TopologicalSpace M₁] [AddCommMonoid M₁] + {M₂ : Type*} [TopologicalSpace M₂] [AddCommMonoid M₂] {M₃ : Type*} [TopologicalSpace M₃] + [AddCommMonoid M₃] {M₄ : Type*} [TopologicalSpace M₄] [AddCommMonoid M₄] [Module R₁ M₁] + [Module R₂ M₂] [Module R₃ M₃] + +/-- A continuous linear equivalence induces a continuous linear map. -/ +@[coe] +def toContinuousLinearMap (e : M₁ ≃SL[σ₁₂] M₂) : M₁ →SL[σ₁₂] M₂ := + { e.toLinearEquiv.toLinearMap with cont := e.continuous_toFun } + +/-- Coerce continuous linear equivs to continuous linear maps. -/ +instance ContinuousLinearMap.coe : Coe (M₁ ≃SL[σ₁₂] M₂) (M₁ →SL[σ₁₂] M₂) := + ⟨toContinuousLinearMap⟩ + +instance equivLike : + EquivLike (M₁ ≃SL[σ₁₂] M₂) M₁ M₂ where + coe f := f.toFun + inv f := f.invFun + coe_injective' f g h₁ h₂ := by + cases' f with f' _ + cases' g with g' _ + rcases f' with ⟨⟨⟨_, _⟩, _⟩, _⟩ + rcases g' with ⟨⟨⟨_, _⟩, _⟩, _⟩ + congr + left_inv f := f.left_inv + right_inv f := f.right_inv + +instance continuousSemilinearEquivClass : + ContinuousSemilinearEquivClass (M₁ ≃SL[σ₁₂] M₂) σ₁₂ M₁ M₂ where + map_add f := f.map_add' + map_smulₛₗ f := f.map_smul' + map_continuous := continuous_toFun + inv_continuous := continuous_invFun + +theorem coe_apply (e : M₁ ≃SL[σ₁₂] M₂) (b : M₁) : (e : M₁ →SL[σ₁₂] M₂) b = e b := + rfl + +@[simp] +theorem coe_toLinearEquiv (f : M₁ ≃SL[σ₁₂] M₂) : ⇑f.toLinearEquiv = f := + rfl + +@[simp, norm_cast] +theorem coe_coe (e : M₁ ≃SL[σ₁₂] M₂) : ⇑(e : M₁ →SL[σ₁₂] M₂) = e := + rfl + +theorem toLinearEquiv_injective : + Function.Injective (toLinearEquiv : (M₁ ≃SL[σ₁₂] M₂) → M₁ ≃ₛₗ[σ₁₂] M₂) := by + rintro ⟨e, _, _⟩ ⟨e', _, _⟩ rfl + rfl + +@[ext] +theorem ext {f g : M₁ ≃SL[σ₁₂] M₂} (h : (f : M₁ → M₂) = g) : f = g := + toLinearEquiv_injective <| LinearEquiv.ext <| congr_fun h + +theorem coe_injective : Function.Injective ((↑) : (M₁ ≃SL[σ₁₂] M₂) → M₁ →SL[σ₁₂] M₂) := + fun _e _e' h => ext <| funext <| ContinuousLinearMap.ext_iff.1 h + +@[simp, norm_cast] +theorem coe_inj {e e' : M₁ ≃SL[σ₁₂] M₂} : (e : M₁ →SL[σ₁₂] M₂) = e' ↔ e = e' := + coe_injective.eq_iff + +/-- A continuous linear equivalence induces a homeomorphism. -/ +def toHomeomorph (e : M₁ ≃SL[σ₁₂] M₂) : M₁ ≃ₜ M₂ := + { e with toEquiv := e.toLinearEquiv.toEquiv } + +@[simp] +theorem coe_toHomeomorph (e : M₁ ≃SL[σ₁₂] M₂) : ⇑e.toHomeomorph = e := + rfl + +theorem isOpenMap (e : M₁ ≃SL[σ₁₂] M₂) : IsOpenMap e := + (ContinuousLinearEquiv.toHomeomorph e).isOpenMap + +theorem image_closure (e : M₁ ≃SL[σ₁₂] M₂) (s : Set M₁) : e '' closure s = closure (e '' s) := + e.toHomeomorph.image_closure s + +theorem preimage_closure (e : M₁ ≃SL[σ₁₂] M₂) (s : Set M₂) : e ⁻¹' closure s = closure (e ⁻¹' s) := + e.toHomeomorph.preimage_closure s + +@[simp] +theorem isClosed_image (e : M₁ ≃SL[σ₁₂] M₂) {s : Set M₁} : IsClosed (e '' s) ↔ IsClosed s := + e.toHomeomorph.isClosed_image + +theorem map_nhds_eq (e : M₁ ≃SL[σ₁₂] M₂) (x : M₁) : map e (𝓝 x) = 𝓝 (e x) := + e.toHomeomorph.map_nhds_eq x + +-- Make some straightforward lemmas available to `simp`. +theorem map_zero (e : M₁ ≃SL[σ₁₂] M₂) : e (0 : M₁) = 0 := + (e : M₁ →SL[σ₁₂] M₂).map_zero + +theorem map_add (e : M₁ ≃SL[σ₁₂] M₂) (x y : M₁) : e (x + y) = e x + e y := + (e : M₁ →SL[σ₁₂] M₂).map_add x y + +@[simp] +theorem map_smulₛₗ (e : M₁ ≃SL[σ₁₂] M₂) (c : R₁) (x : M₁) : e (c • x) = σ₁₂ c • e x := + (e : M₁ →SL[σ₁₂] M₂).map_smulₛₗ c x + +theorem map_smul [Module R₁ M₂] (e : M₁ ≃L[R₁] M₂) (c : R₁) (x : M₁) : e (c • x) = c • e x := + (e : M₁ →L[R₁] M₂).map_smul c x + +theorem map_eq_zero_iff (e : M₁ ≃SL[σ₁₂] M₂) {x : M₁} : e x = 0 ↔ x = 0 := + e.toLinearEquiv.map_eq_zero_iff + +attribute [continuity] + ContinuousLinearEquiv.continuous_toFun ContinuousLinearEquiv.continuous_invFun + +@[continuity] +protected theorem continuous (e : M₁ ≃SL[σ₁₂] M₂) : Continuous (e : M₁ → M₂) := + e.continuous_toFun + +protected theorem continuousOn (e : M₁ ≃SL[σ₁₂] M₂) {s : Set M₁} : ContinuousOn (e : M₁ → M₂) s := + e.continuous.continuousOn + +protected theorem continuousAt (e : M₁ ≃SL[σ₁₂] M₂) {x : M₁} : ContinuousAt (e : M₁ → M₂) x := + e.continuous.continuousAt + +protected theorem continuousWithinAt (e : M₁ ≃SL[σ₁₂] M₂) {s : Set M₁} {x : M₁} : + ContinuousWithinAt (e : M₁ → M₂) s x := + e.continuous.continuousWithinAt + +theorem comp_continuousOn_iff {α : Type*} [TopologicalSpace α] (e : M₁ ≃SL[σ₁₂] M₂) {f : α → M₁} + {s : Set α} : ContinuousOn (e ∘ f) s ↔ ContinuousOn f s := + e.toHomeomorph.comp_continuousOn_iff _ _ + +theorem comp_continuous_iff {α : Type*} [TopologicalSpace α] (e : M₁ ≃SL[σ₁₂] M₂) {f : α → M₁} : + Continuous (e ∘ f) ↔ Continuous f := + e.toHomeomorph.comp_continuous_iff + +/-- An extensionality lemma for `R ≃L[R] M`. -/ +theorem ext₁ [TopologicalSpace R₁] {f g : R₁ ≃L[R₁] M₁} (h : f 1 = g 1) : f = g := + ext <| funext fun x => mul_one x ▸ by rw [← smul_eq_mul, map_smul, h, map_smul] + +section + +variable (R₁ M₁) + +/-- The identity map as a continuous linear equivalence. -/ +@[refl] +protected def refl : M₁ ≃L[R₁] M₁ := + { LinearEquiv.refl R₁ M₁ with + continuous_toFun := continuous_id + continuous_invFun := continuous_id } + +@[simp] +theorem refl_apply (x : M₁) : + ContinuousLinearEquiv.refl R₁ M₁ x = x := rfl + +end + +@[simp, norm_cast] +theorem coe_refl : ↑(ContinuousLinearEquiv.refl R₁ M₁) = ContinuousLinearMap.id R₁ M₁ := + rfl + +@[simp, norm_cast] +theorem coe_refl' : ⇑(ContinuousLinearEquiv.refl R₁ M₁) = id := + rfl + +/-- The inverse of a continuous linear equivalence as a continuous linear equivalence -/ +@[symm] +protected def symm (e : M₁ ≃SL[σ₁₂] M₂) : M₂ ≃SL[σ₂₁] M₁ := + { e.toLinearEquiv.symm with + continuous_toFun := e.continuous_invFun + continuous_invFun := e.continuous_toFun } + +@[simp] +theorem symm_toLinearEquiv (e : M₁ ≃SL[σ₁₂] M₂) : e.symm.toLinearEquiv = e.toLinearEquiv.symm := by + ext + rfl + +@[simp] +theorem symm_toHomeomorph (e : M₁ ≃SL[σ₁₂] M₂) : e.toHomeomorph.symm = e.symm.toHomeomorph := + rfl + +/-- See Note [custom simps projection]. We need to specify this projection explicitly in this case, + because it is a composition of multiple projections. -/ +def Simps.apply (h : M₁ ≃SL[σ₁₂] M₂) : M₁ → M₂ := + h + +/-- See Note [custom simps projection] -/ +def Simps.symm_apply (h : M₁ ≃SL[σ₁₂] M₂) : M₂ → M₁ := + h.symm + +initialize_simps_projections ContinuousLinearEquiv (toFun → apply, invFun → symm_apply) + +theorem symm_map_nhds_eq (e : M₁ ≃SL[σ₁₂] M₂) (x : M₁) : map e.symm (𝓝 (e x)) = 𝓝 x := + e.toHomeomorph.symm_map_nhds_eq x + +/-- The composition of two continuous linear equivalences as a continuous linear equivalence. -/ +@[trans] +protected def trans (e₁ : M₁ ≃SL[σ₁₂] M₂) (e₂ : M₂ ≃SL[σ₂₃] M₃) : M₁ ≃SL[σ₁₃] M₃ := + { e₁.toLinearEquiv.trans e₂.toLinearEquiv with + continuous_toFun := e₂.continuous_toFun.comp e₁.continuous_toFun + continuous_invFun := e₁.continuous_invFun.comp e₂.continuous_invFun } + +@[simp] +theorem trans_toLinearEquiv (e₁ : M₁ ≃SL[σ₁₂] M₂) (e₂ : M₂ ≃SL[σ₂₃] M₃) : + (e₁.trans e₂).toLinearEquiv = e₁.toLinearEquiv.trans e₂.toLinearEquiv := by + ext + rfl + +/-- Product of two continuous linear equivalences. The map comes from `Equiv.prodCongr`. -/ +def prod [Module R₁ M₂] [Module R₁ M₃] [Module R₁ M₄] (e : M₁ ≃L[R₁] M₂) (e' : M₃ ≃L[R₁] M₄) : + (M₁ × M₃) ≃L[R₁] M₂ × M₄ := + { e.toLinearEquiv.prod e'.toLinearEquiv with + continuous_toFun := e.continuous_toFun.prodMap e'.continuous_toFun + continuous_invFun := e.continuous_invFun.prodMap e'.continuous_invFun } + +@[simp, norm_cast] +theorem prod_apply [Module R₁ M₂] [Module R₁ M₃] [Module R₁ M₄] (e : M₁ ≃L[R₁] M₂) + (e' : M₃ ≃L[R₁] M₄) (x) : e.prod e' x = (e x.1, e' x.2) := + rfl + +@[simp, norm_cast] +theorem coe_prod [Module R₁ M₂] [Module R₁ M₃] [Module R₁ M₄] (e : M₁ ≃L[R₁] M₂) + (e' : M₃ ≃L[R₁] M₄) : + (e.prod e' : M₁ × M₃ →L[R₁] M₂ × M₄) = (e : M₁ →L[R₁] M₂).prodMap (e' : M₃ →L[R₁] M₄) := + rfl + +theorem prod_symm [Module R₁ M₂] [Module R₁ M₃] [Module R₁ M₄] (e : M₁ ≃L[R₁] M₂) + (e' : M₃ ≃L[R₁] M₄) : (e.prod e').symm = e.symm.prod e'.symm := + rfl + +variable (R₁ M₁ M₂) + +/-- Product of modules is commutative up to continuous linear isomorphism. -/ +@[simps! apply toLinearEquiv] +def prodComm [Module R₁ M₂] : (M₁ × M₂) ≃L[R₁] M₂ × M₁ := + { LinearEquiv.prodComm R₁ M₁ M₂ with + continuous_toFun := continuous_swap + continuous_invFun := continuous_swap } + +@[simp] lemma prodComm_symm [Module R₁ M₂] : (prodComm R₁ M₁ M₂).symm = prodComm R₁ M₂ M₁ := rfl + +variable {R₁ M₁ M₂} + +protected theorem bijective (e : M₁ ≃SL[σ₁₂] M₂) : Function.Bijective e := + e.toLinearEquiv.toEquiv.bijective + +protected theorem injective (e : M₁ ≃SL[σ₁₂] M₂) : Function.Injective e := + e.toLinearEquiv.toEquiv.injective + +protected theorem surjective (e : M₁ ≃SL[σ₁₂] M₂) : Function.Surjective e := + e.toLinearEquiv.toEquiv.surjective + +@[simp] +theorem trans_apply (e₁ : M₁ ≃SL[σ₁₂] M₂) (e₂ : M₂ ≃SL[σ₂₃] M₃) (c : M₁) : + (e₁.trans e₂) c = e₂ (e₁ c) := + rfl + +@[simp] +theorem apply_symm_apply (e : M₁ ≃SL[σ₁₂] M₂) (c : M₂) : e (e.symm c) = c := + e.1.right_inv c + +@[simp] +theorem symm_apply_apply (e : M₁ ≃SL[σ₁₂] M₂) (b : M₁) : e.symm (e b) = b := + e.1.left_inv b + +@[simp] +theorem symm_trans_apply (e₁ : M₂ ≃SL[σ₂₁] M₁) (e₂ : M₃ ≃SL[σ₃₂] M₂) (c : M₁) : + (e₂.trans e₁).symm c = e₂.symm (e₁.symm c) := + rfl + +@[simp] +theorem symm_image_image (e : M₁ ≃SL[σ₁₂] M₂) (s : Set M₁) : e.symm '' (e '' s) = s := + e.toLinearEquiv.toEquiv.symm_image_image s + +@[simp] +theorem image_symm_image (e : M₁ ≃SL[σ₁₂] M₂) (s : Set M₂) : e '' (e.symm '' s) = s := + e.symm.symm_image_image s + +@[simp, norm_cast] +theorem comp_coe (f : M₁ ≃SL[σ₁₂] M₂) (f' : M₂ ≃SL[σ₂₃] M₃) : + (f' : M₂ →SL[σ₂₃] M₃).comp (f : M₁ →SL[σ₁₂] M₂) = (f.trans f' : M₁ →SL[σ₁₃] M₃) := + rfl + +-- Porting note: The priority should be higher than `comp_coe`. +@[simp high] +theorem coe_comp_coe_symm (e : M₁ ≃SL[σ₁₂] M₂) : + (e : M₁ →SL[σ₁₂] M₂).comp (e.symm : M₂ →SL[σ₂₁] M₁) = ContinuousLinearMap.id R₂ M₂ := + ContinuousLinearMap.ext e.apply_symm_apply + +-- Porting note: The priority should be higher than `comp_coe`. +@[simp high] +theorem coe_symm_comp_coe (e : M₁ ≃SL[σ₁₂] M₂) : + (e.symm : M₂ →SL[σ₂₁] M₁).comp (e : M₁ →SL[σ₁₂] M₂) = ContinuousLinearMap.id R₁ M₁ := + ContinuousLinearMap.ext e.symm_apply_apply + +@[simp] +theorem symm_comp_self (e : M₁ ≃SL[σ₁₂] M₂) : (e.symm : M₂ → M₁) ∘ (e : M₁ → M₂) = id := by + ext x + exact symm_apply_apply e x + +@[simp] +theorem self_comp_symm (e : M₁ ≃SL[σ₁₂] M₂) : (e : M₁ → M₂) ∘ (e.symm : M₂ → M₁) = id := by + ext x + exact apply_symm_apply e x + +@[simp] +theorem symm_symm (e : M₁ ≃SL[σ₁₂] M₂) : e.symm.symm = e := rfl + +@[simp] +theorem refl_symm : (ContinuousLinearEquiv.refl R₁ M₁).symm = ContinuousLinearEquiv.refl R₁ M₁ := + rfl + +theorem symm_symm_apply (e : M₁ ≃SL[σ₁₂] M₂) (x : M₁) : e.symm.symm x = e x := + rfl + +theorem symm_apply_eq (e : M₁ ≃SL[σ₁₂] M₂) {x y} : e.symm x = y ↔ x = e y := + e.toLinearEquiv.symm_apply_eq + +theorem eq_symm_apply (e : M₁ ≃SL[σ₁₂] M₂) {x y} : y = e.symm x ↔ e y = x := + e.toLinearEquiv.eq_symm_apply + +protected theorem image_eq_preimage (e : M₁ ≃SL[σ₁₂] M₂) (s : Set M₁) : e '' s = e.symm ⁻¹' s := + e.toLinearEquiv.toEquiv.image_eq_preimage s + +protected theorem image_symm_eq_preimage (e : M₁ ≃SL[σ₁₂] M₂) (s : Set M₂) : + e.symm '' s = e ⁻¹' s := by rw [e.symm.image_eq_preimage, e.symm_symm] + +@[simp] +protected theorem symm_preimage_preimage (e : M₁ ≃SL[σ₁₂] M₂) (s : Set M₂) : + e.symm ⁻¹' (e ⁻¹' s) = s := + e.toLinearEquiv.toEquiv.symm_preimage_preimage s + +@[simp] +protected theorem preimage_symm_preimage (e : M₁ ≃SL[σ₁₂] M₂) (s : Set M₁) : + e ⁻¹' (e.symm ⁻¹' s) = s := + e.symm.symm_preimage_preimage s + +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₂) : IsUniformEmbedding e := + e.toLinearEquiv.toEquiv.isUniformEmbedding e.toContinuousLinearMap.uniformContinuous + e.symm.toContinuousLinearMap.uniformContinuous + +@[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) : 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. See also `equivOfInverse'`. -/ +def equivOfInverse (f₁ : M₁ →SL[σ₁₂] M₂) (f₂ : M₂ →SL[σ₂₁] M₁) (h₁ : Function.LeftInverse f₂ f₁) + (h₂ : Function.RightInverse f₂ f₁) : M₁ ≃SL[σ₁₂] M₂ := + { f₁ with + continuous_toFun := f₁.continuous + invFun := f₂ + continuous_invFun := f₂.continuous + left_inv := h₁ + right_inv := h₂ } + +@[simp] +theorem equivOfInverse_apply (f₁ : M₁ →SL[σ₁₂] M₂) (f₂ h₁ h₂ x) : + equivOfInverse f₁ f₂ h₁ h₂ x = f₁ x := + rfl + +@[simp] +theorem symm_equivOfInverse (f₁ : M₁ →SL[σ₁₂] M₂) (f₂ h₁ h₂) : + (equivOfInverse f₁ f₂ h₁ h₂).symm = equivOfInverse f₂ f₁ h₂ h₁ := + rfl + +/-- Create a `ContinuousLinearEquiv` from two `ContinuousLinearMap`s that are +inverse of each other, in the `ContinuousLinearMap.comp` sense. See also `equivOfInverse`. -/ +def equivOfInverse' (f₁ : M₁ →SL[σ₁₂] M₂) (f₂ : M₂ →SL[σ₂₁] M₁) + (h₁ : f₁.comp f₂ = .id R₂ M₂) (h₂ : f₂.comp f₁ = .id R₁ M₁) : M₁ ≃SL[σ₁₂] M₂ := + equivOfInverse f₁ f₂ + (fun x ↦ by simpa using congr($(h₂) x)) (fun x ↦ by simpa using congr($(h₁) x)) + +@[simp] +theorem equivOfInverse'_apply (f₁ : M₁ →SL[σ₁₂] M₂) (f₂ h₁ h₂ x) : + equivOfInverse' f₁ f₂ h₁ h₂ x = f₁ x := + rfl + +/-- The inverse of `equivOfInverse'` is obtained by swapping the order of its parameters. -/ +@[simp] +theorem symm_equivOfInverse' (f₁ : M₁ →SL[σ₁₂] M₂) (f₂ h₁ h₂) : + (equivOfInverse' f₁ f₂ h₁ h₂).symm = equivOfInverse' f₂ f₁ h₂ h₁ := + rfl + +variable (M₁) + +/-- The continuous linear equivalences from `M` to itself form a group under composition. -/ +instance automorphismGroup : Group (M₁ ≃L[R₁] M₁) where + mul f g := g.trans f + one := ContinuousLinearEquiv.refl R₁ M₁ + inv f := f.symm + mul_assoc f g h := by + ext + rfl + mul_one f := by + ext + rfl + one_mul f := by + ext + rfl + inv_mul_cancel f := by + ext x + exact f.left_inv x + +variable {M₁} {R₄ : Type*} [Semiring R₄] [Module R₄ M₄] {σ₃₄ : R₃ →+* R₄} {σ₄₃ : R₄ →+* R₃} + [RingHomInvPair σ₃₄ σ₄₃] [RingHomInvPair σ₄₃ σ₃₄] {σ₂₄ : R₂ →+* R₄} {σ₁₄ : R₁ →+* R₄} + [RingHomCompTriple σ₂₁ σ₁₄ σ₂₄] [RingHomCompTriple σ₂₄ σ₄₃ σ₂₃] [RingHomCompTriple σ₁₃ σ₃₄ σ₁₄] + +/-- The continuous linear equivalence between `ULift M₁` and `M₁`. + +This is a continuous version of `ULift.moduleEquiv`. -/ +def ulift : ULift M₁ ≃L[R₁] M₁ := + { ULift.moduleEquiv with + continuous_toFun := continuous_uLift_down + continuous_invFun := continuous_uLift_up } + +/-- A pair of continuous (semi)linear equivalences generates an equivalence between the spaces of +continuous linear maps. See also `ContinuousLinearEquiv.arrowCongr`. -/ +@[simps] +def arrowCongrEquiv (e₁₂ : M₁ ≃SL[σ₁₂] M₂) (e₄₃ : M₄ ≃SL[σ₄₃] M₃) : + (M₁ →SL[σ₁₄] M₄) ≃ (M₂ →SL[σ₂₃] M₃) where + toFun f := (e₄₃ : M₄ →SL[σ₄₃] M₃).comp (f.comp (e₁₂.symm : M₂ →SL[σ₂₁] M₁)) + invFun f := (e₄₃.symm : M₃ →SL[σ₃₄] M₄).comp (f.comp (e₁₂ : M₁ →SL[σ₁₂] M₂)) + left_inv f := + ContinuousLinearMap.ext fun x => by + simp only [ContinuousLinearMap.comp_apply, symm_apply_apply, coe_coe] + right_inv f := + ContinuousLinearMap.ext fun x => by + simp only [ContinuousLinearMap.comp_apply, apply_symm_apply, coe_coe] + +section piCongrRight + +variable {ι : Type*} {M : ι → Type*} [∀ i, TopologicalSpace (M i)] [∀ i, AddCommMonoid (M i)] + [∀ i, Module R₁ (M i)] {N : ι → Type*} [∀ i, TopologicalSpace (N i)] [∀ i, AddCommMonoid (N i)] + [∀ i, Module R₁ (N i)] (f : (i : ι) → M i ≃L[R₁] N i) + +/-- Combine a family of continuous linear equivalences into a continuous linear equivalence of +pi-types. -/ +def piCongrRight : ((i : ι) → M i) ≃L[R₁] (i : ι) → N i := + { LinearEquiv.piCongrRight fun i ↦ f i with + continuous_toFun := by + exact continuous_pi fun i ↦ (f i).continuous_toFun.comp (continuous_apply i) + continuous_invFun := by + exact continuous_pi fun i => (f i).continuous_invFun.comp (continuous_apply i) } + +@[simp] +theorem piCongrRight_apply (m : (i : ι) → M i) (i : ι) : + piCongrRight f m i = (f i) (m i) := rfl + +@[simp] +theorem piCongrRight_symm_apply (n : (i : ι) → N i) (i : ι) : + (piCongrRight f).symm n i = (f i).symm (n i) := rfl + +end piCongrRight + +end AddCommMonoid + +section AddCommGroup + +variable {R : Type*} [Semiring R] {M : Type*} [TopologicalSpace M] [AddCommGroup M] {M₂ : Type*} + [TopologicalSpace M₂] [AddCommGroup M₂] {M₃ : Type*} [TopologicalSpace M₃] [AddCommGroup M₃] + {M₄ : Type*} [TopologicalSpace M₄] [AddCommGroup M₄] [Module R M] [Module R M₂] [Module R M₃] + [Module R M₄] + +variable [TopologicalAddGroup M₄] + +/-- Equivalence given by a block lower diagonal matrix. `e` and `e'` are diagonal square blocks, + and `f` is a rectangular block below the diagonal. -/ +def skewProd (e : M ≃L[R] M₂) (e' : M₃ ≃L[R] M₄) (f : M →L[R] M₄) : (M × M₃) ≃L[R] M₂ × M₄ := + { + e.toLinearEquiv.skewProd e'.toLinearEquiv + ↑f with + continuous_toFun := + (e.continuous_toFun.comp continuous_fst).prod_mk + ((e'.continuous_toFun.comp continuous_snd).add <| f.continuous.comp continuous_fst) + continuous_invFun := + (e.continuous_invFun.comp continuous_fst).prod_mk + (e'.continuous_invFun.comp <| + continuous_snd.sub <| f.continuous.comp <| e.continuous_invFun.comp continuous_fst) } + +@[simp] +theorem skewProd_apply (e : M ≃L[R] M₂) (e' : M₃ ≃L[R] M₄) (f : M →L[R] M₄) (x) : + e.skewProd e' f x = (e x.1, e' x.2 + f x.1) := + rfl + +@[simp] +theorem skewProd_symm_apply (e : M ≃L[R] M₂) (e' : M₃ ≃L[R] M₄) (f : M →L[R] M₄) (x) : + (e.skewProd e' f).symm x = (e.symm x.1, e'.symm (x.2 - f (e.symm x.1))) := + rfl + +variable (R) in +/-- The negation map as a continuous linear equivalence. -/ +def neg [ContinuousNeg M] : + M ≃L[R] M := + { LinearEquiv.neg R with + continuous_toFun := continuous_neg + continuous_invFun := continuous_neg } + +@[simp] +theorem coe_neg [ContinuousNeg M] : + (neg R : M → M) = -id := rfl + +@[simp] +theorem neg_apply [ContinuousNeg M] (x : M) : + neg R x = -x := by simp + +@[simp] +theorem symm_neg [ContinuousNeg M] : + (neg R : M ≃L[R] M).symm = neg R := rfl + +end AddCommGroup + +section Ring + +variable {R : Type*} [Ring R] {R₂ : Type*} [Ring R₂] {M : Type*} [TopologicalSpace M] + [AddCommGroup M] [Module R M] {M₂ : Type*} [TopologicalSpace M₂] [AddCommGroup M₂] [Module R₂ M₂] + +variable {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] + +theorem map_sub (e : M ≃SL[σ₁₂] M₂) (x y : M) : e (x - y) = e x - e y := + (e : M →SL[σ₁₂] M₂).map_sub x y + +theorem map_neg (e : M ≃SL[σ₁₂] M₂) (x : M) : e (-x) = -e x := + (e : M →SL[σ₁₂] M₂).map_neg x + +section + +/-! The next theorems cover the identification between `M ≃L[R] M`and the group of units of the ring +`M →L[R] M`. -/ + +/-- An invertible continuous linear map `f` determines a continuous equivalence from `M` to itself. +-/ +def ofUnit (f : (M →L[R] M)ˣ) : M ≃L[R] M where + toLinearEquiv := + { toFun := f.val + map_add' := by simp + map_smul' := by simp + invFun := f.inv + left_inv := fun x => + show (f.inv * f.val) x = x by + rw [f.inv_val] + simp + right_inv := fun x => + show (f.val * f.inv) x = x by + rw [f.val_inv] + simp } + continuous_toFun := f.val.continuous + continuous_invFun := f.inv.continuous + +/-- A continuous equivalence from `M` to itself determines an invertible continuous linear map. -/ +def toUnit (f : M ≃L[R] M) : (M →L[R] M)ˣ where + val := f + inv := f.symm + val_inv := by + ext + simp + inv_val := by + ext + simp + +variable (R M) + +/-- The units of the algebra of continuous `R`-linear endomorphisms of `M` is multiplicatively +equivalent to the type of continuous linear equivalences between `M` and itself. -/ +def unitsEquiv : (M →L[R] M)ˣ ≃* M ≃L[R] M where + toFun := ofUnit + invFun := toUnit + left_inv f := by + ext + rfl + right_inv f := by + ext + rfl + map_mul' x y := by + ext + rfl + +@[simp] +theorem unitsEquiv_apply (f : (M →L[R] M)ˣ) (x : M) : unitsEquiv R M f x = (f : M →L[R] M) x := + rfl + +end + +section + +variable (R) [TopologicalSpace R] +variable [ContinuousMul R] + +/-- Continuous linear equivalences `R ≃L[R] R` are enumerated by `Rˣ`. -/ +def unitsEquivAut : Rˣ ≃ R ≃L[R] R where + toFun u := + equivOfInverse (ContinuousLinearMap.smulRight (1 : R →L[R] R) ↑u) + (ContinuousLinearMap.smulRight (1 : R →L[R] R) ↑u⁻¹) (fun x => by simp) fun x => by simp + invFun e := + ⟨e 1, e.symm 1, by rw [← smul_eq_mul, ← map_smul, smul_eq_mul, mul_one, symm_apply_apply], by + rw [← smul_eq_mul, ← map_smul, smul_eq_mul, mul_one, apply_symm_apply]⟩ + left_inv u := Units.ext <| by simp + right_inv e := ext₁ <| by simp + +variable {R} + +@[simp] +theorem unitsEquivAut_apply (u : Rˣ) (x : R) : unitsEquivAut R u x = x * u := + rfl + +@[simp] +theorem unitsEquivAut_apply_symm (u : Rˣ) (x : R) : (unitsEquivAut R u).symm x = x * ↑u⁻¹ := + rfl + +@[simp] +theorem unitsEquivAut_symm_apply (e : R ≃L[R] R) : ↑((unitsEquivAut R).symm e) = e 1 := + rfl + +end + +variable [Module R M₂] [TopologicalAddGroup M] + +/-- A pair of continuous linear maps such that `f₁ ∘ f₂ = id` generates a continuous +linear equivalence `e` between `M` and `M₂ × f₁.ker` such that `(e x).2 = x` for `x ∈ f₁.ker`, +`(e x).1 = f₁ x`, and `(e (f₂ y)).2 = 0`. The map is given by `e x = (f₁ x, x - f₂ (f₁ x))`. -/ +def equivOfRightInverse (f₁ : M →L[R] M₂) (f₂ : M₂ →L[R] M) (h : Function.RightInverse f₂ f₁) : + M ≃L[R] M₂ × ker f₁ := + equivOfInverse (f₁.prod (f₁.projKerOfRightInverse f₂ h)) (f₂.coprod (ker f₁).subtypeL) + (fun x => by simp) fun ⟨x, y⟩ => by + -- Porting note: `simp` timeouts. + rw [ContinuousLinearMap.coprod_apply, + Submodule.subtypeL_apply, _root_.map_add, ContinuousLinearMap.prod_apply, h x, + ContinuousLinearMap.projKerOfRightInverse_comp_inv, + ContinuousLinearMap.prod_apply, LinearMap.map_coe_ker, + ContinuousLinearMap.projKerOfRightInverse_apply_idem, Prod.mk_add_mk, add_zero, zero_add] + +@[simp] +theorem fst_equivOfRightInverse (f₁ : M →L[R] M₂) (f₂ : M₂ →L[R] M) + (h : Function.RightInverse f₂ f₁) (x : M) : (equivOfRightInverse f₁ f₂ h x).1 = f₁ x := + rfl + +@[simp] +theorem snd_equivOfRightInverse (f₁ : M →L[R] M₂) (f₂ : M₂ →L[R] M) + (h : Function.RightInverse f₂ f₁) (x : M) : + ((equivOfRightInverse f₁ f₂ h x).2 : M) = x - f₂ (f₁ x) := + rfl + +@[simp] +theorem equivOfRightInverse_symm_apply (f₁ : M →L[R] M₂) (f₂ : M₂ →L[R] M) + (h : Function.RightInverse f₂ f₁) (y : M₂ × ker f₁) : + (equivOfRightInverse f₁ f₂ h).symm y = f₂ y.1 + y.2 := + rfl + +end Ring + +section + +variable (ι R M : Type*) [Unique ι] [Semiring R] [AddCommMonoid M] [Module R M] + [TopologicalSpace M] + +/-- If `ι` has a unique element, then `ι → M` is continuously linear equivalent to `M`. -/ +def funUnique : (ι → M) ≃L[R] M := + { Homeomorph.funUnique ι M with toLinearEquiv := LinearEquiv.funUnique ι R M } + +variable {ι R M} + +@[simp] +theorem coe_funUnique : ⇑(funUnique ι R M) = Function.eval default := + rfl + +@[simp] +theorem coe_funUnique_symm : ⇑(funUnique ι R M).symm = Function.const ι := + rfl + +variable (R M) + +/-- Continuous linear equivalence between dependent functions `(i : Fin 2) → M i` and `M 0 × M 1`. +-/ +@[simps! (config := .asFn) apply symm_apply] +def piFinTwo (M : Fin 2 → Type*) [∀ i, AddCommMonoid (M i)] [∀ i, Module R (M i)] + [∀ i, TopologicalSpace (M i)] : ((i : _) → M i) ≃L[R] M 0 × M 1 := + { Homeomorph.piFinTwo M with toLinearEquiv := LinearEquiv.piFinTwo R M } + +/-- Continuous linear equivalence between vectors in `M² = Fin 2 → M` and `M × M`. -/ +@[simps! (config := .asFn) apply symm_apply] +def finTwoArrow : (Fin 2 → M) ≃L[R] M × M := + { piFinTwo R fun _ => M with toLinearEquiv := LinearEquiv.finTwoArrow R M } + +end + +end ContinuousLinearEquiv + +namespace ContinuousLinearMap + +variable {R : Type*} {M M₂ M₃ : Type*} + [TopologicalSpace M] [TopologicalSpace M₂] [TopologicalSpace M₃] + +section + +variable [Semiring R] + [AddCommMonoid M] [Module R M] + [AddCommMonoid M₂] [Module R M₂] + [AddCommMonoid M₃] [Module R M₃] + +/-- A continuous linear map is invertible if it is the forward direction of a continuous linear +equivalence. -/ +def IsInvertible (f : M →L[R] M₂) : Prop := + ∃ (A : M ≃L[R] M₂), A = f + +open Classical in +/-- Introduce a function `inverse` from `M →L[R] M₂` to `M₂ →L[R] M`, which sends `f` to `f.symm` if +`f` is a continuous linear equivalence and to `0` otherwise. This definition is somewhat ad hoc, +but one needs a fully (rather than partially) defined inverse function for some purposes, including +for calculus. -/ +noncomputable def inverse : (M →L[R] M₂) → M₂ →L[R] M := fun f => + if h : f.IsInvertible then ((Classical.choose h).symm : M₂ →L[R] M) else 0 + +@[simp] lemma isInvertible_equiv {f : M ≃L[R] M₂} : IsInvertible (f : M →L[R] M₂) := ⟨f, rfl⟩ + +/-- By definition, if `f` is invertible then `inverse f = f.symm`. -/ +@[simp] +theorem inverse_equiv (e : M ≃L[R] M₂) : inverse (e : M →L[R] M₂) = e.symm := by + simp [inverse] + +/-- By definition, if `f` is not invertible then `inverse f = 0`. -/ +@[simp] lemma inverse_of_not_isInvertible + {f : M →L[R] M₂} (hf : ¬ f.IsInvertible) : f.inverse = 0 := + dif_neg hf + +@[deprecated (since := "2024-10-29")] alias inverse_non_equiv := inverse_of_not_isInvertible + +@[simp] theorem inverse_zero : inverse (0 : M →L[R] M₂) = 0 := by + by_cases h : IsInvertible (0 : M →L[R] M₂) + · rcases h with ⟨e', he'⟩ + simp only [← he', inverse_equiv] + ext v + apply e'.injective + rw [← ContinuousLinearEquiv.coe_coe, he'] + rfl + · exact inverse_of_not_isInvertible h + +lemma IsInvertible.comp {g : M₂ →L[R] M₃} {f : M →L[R] M₂} + (hg : g.IsInvertible) (hf : f.IsInvertible) : (g ∘L f).IsInvertible := by + rcases hg with ⟨N, rfl⟩ + rcases hf with ⟨M, rfl⟩ + exact ⟨M.trans N, rfl⟩ + +lemma IsInvertible.of_inverse {f : M →L[R] M₂} {g : M₂ →L[R] M} + (hf : f ∘L g = id R M₂) (hg : g ∘L f = id R M) : + f.IsInvertible := + ⟨ContinuousLinearEquiv.equivOfInverse' _ _ hf hg, rfl⟩ + +lemma inverse_eq {f : M →L[R] M₂} {g : M₂ →L[R] M} (hf : f ∘L g = id R M₂) (hg : g ∘L f = id R M) : + f.inverse = g := by + have : f = ContinuousLinearEquiv.equivOfInverse' f g hf hg := rfl + rw [this, inverse_equiv] + rfl + +lemma IsInvertible.inverse_apply_eq {f : M →L[R] M₂} {x : M} {y : M₂} (hf : f.IsInvertible) : + f.inverse y = x ↔ y = f x := by + rcases hf with ⟨M, rfl⟩ + simp only [inverse_equiv, ContinuousLinearEquiv.coe_coe] + exact ContinuousLinearEquiv.symm_apply_eq M + +@[simp] lemma isInvertible_equiv_comp {e : M₂ ≃L[R] M₃} {f : M →L[R] M₂} : + ((e : M₂ →L[R] M₃) ∘L f).IsInvertible ↔ f.IsInvertible := by + constructor + · rintro ⟨A, hA⟩ + have : f = e.symm ∘L ((e : M₂ →L[R] M₃) ∘L f) := by ext; simp + rw [this, ← hA] + simp + · rintro ⟨M, rfl⟩ + simp + +@[simp] lemma isInvertible_comp_equiv {e : M₃ ≃L[R] M} {f : M →L[R] M₂} : + (f ∘L (e : M₃ →L[R] M)).IsInvertible ↔ f.IsInvertible := by + constructor + · rintro ⟨A, hA⟩ + have : f = (f ∘L (e : M₃ →L[R] M)) ∘L e.symm := by ext; simp + rw [this, ← hA] + simp + · rintro ⟨M, rfl⟩ + simp + +@[simp] lemma inverse_equiv_comp {e : M₂ ≃L[R] M₃} {f : M →L[R] M₂} : + (e ∘L f).inverse = f.inverse ∘L (e.symm : M₃ →L[R] M₂) := by + by_cases hf : f.IsInvertible + · rcases hf with ⟨A, rfl⟩ + simp only [ContinuousLinearEquiv.comp_coe, inverse_equiv, ContinuousLinearEquiv.coe_inj] + rfl + · rw [inverse_of_not_isInvertible (by simp [hf]), inverse_of_not_isInvertible hf] + rfl + +@[simp] lemma inverse_comp_equiv {e : M₃ ≃L[R] M} {f : M →L[R] M₂} : + (f ∘L e).inverse = (e.symm : M →L[R] M₃) ∘L f.inverse := by + by_cases hf : f.IsInvertible + · rcases hf with ⟨A, rfl⟩ + simp only [ContinuousLinearEquiv.comp_coe, inverse_equiv, ContinuousLinearEquiv.coe_inj] + rfl + · rw [inverse_of_not_isInvertible (by simp [hf]), inverse_of_not_isInvertible hf] + simp + +lemma IsInvertible.inverse_comp_of_left {g : M₂ →L[R] M₃} {f : M →L[R] M₂} + (hg : g.IsInvertible) : (g ∘L f).inverse = f.inverse ∘L g.inverse := by + rcases hg with ⟨N, rfl⟩ + simp + +lemma IsInvertible.inverse_comp_apply_of_left {g : M₂ →L[R] M₃} {f : M →L[R] M₂} {v : M₃} + (hg : g.IsInvertible) : (g ∘L f).inverse v = f.inverse (g.inverse v) := by + simp only [hg.inverse_comp_of_left, coe_comp', Function.comp_apply] + +lemma IsInvertible.inverse_comp_of_right {g : M₂ →L[R] M₃} {f : M →L[R] M₂} + (hf : f.IsInvertible) : (g ∘L f).inverse = f.inverse ∘L g.inverse := by + rcases hf with ⟨M, rfl⟩ + simp + +lemma IsInvertible.inverse_comp_apply_of_right {g : M₂ →L[R] M₃} {f : M →L[R] M₂} {v : M₃} + (hf : f.IsInvertible) : (g ∘L f).inverse v = f.inverse (g.inverse v) := by + simp only [hf.inverse_comp_of_right, coe_comp', Function.comp_apply] + +end + +section + +variable [Ring R] +variable [AddCommGroup M] [Module R M] +variable [AddCommGroup M₂] [Module R M₂] + +@[simp] +theorem ring_inverse_equiv (e : M ≃L[R] M) : Ring.inverse ↑e = inverse (e : M →L[R] M) := by + suffices Ring.inverse ((ContinuousLinearEquiv.unitsEquiv _ _).symm e : M →L[R] M) = inverse ↑e by + convert this + simp + rfl + +/-- The function `ContinuousLinearEquiv.inverse` can be written in terms of `Ring.inverse` for the +ring of self-maps of the domain. -/ +theorem to_ring_inverse (e : M ≃L[R] M₂) (f : M →L[R] M₂) : + inverse f = Ring.inverse ((e.symm : M₂ →L[R] M).comp f) ∘L e.symm := by + by_cases h₁ : f.IsInvertible + · obtain ⟨e', he'⟩ := h₁ + rw [← he'] + change _ = Ring.inverse (e'.trans e.symm : M →L[R] M) ∘L (e.symm : M₂ →L[R] M) + ext + simp + · suffices ¬IsUnit ((e.symm : M₂ →L[R] M).comp f) by simp [this, h₁] + contrapose! h₁ + rcases h₁ with ⟨F, hF⟩ + use (ContinuousLinearEquiv.unitsEquiv _ _ F).trans e + ext + dsimp + rw [hF] + simp + +theorem ring_inverse_eq_map_inverse : Ring.inverse = @inverse R M M _ _ _ _ _ _ _ := by + ext + simp [to_ring_inverse (ContinuousLinearEquiv.refl R M)] + +@[simp] theorem inverse_id : (id R M).inverse = id R M := by + rw [← ring_inverse_eq_map_inverse] + exact Ring.inverse_one _ + +end + +end ContinuousLinearMap + +namespace Submodule + +variable {R : Type*} [Ring R] {M : Type*} [TopologicalSpace M] [AddCommGroup M] [Module R M] + +open ContinuousLinearMap + +/-- If `p` is a closed complemented submodule, +then there exists a submodule `q` and a continuous linear equivalence `M ≃L[R] (p × q)` such that +`e (x : p) = (x, 0)`, `e (y : q) = (0, y)`, and `e.symm x = x.1 + x.2`. + +In fact, the properties of `e` imply the properties of `e.symm` and vice versa, +but we provide both for convenience. -/ +lemma ClosedComplemented.exists_submodule_equiv_prod [TopologicalAddGroup M] + {p : Submodule R M} (hp : p.ClosedComplemented) : + ∃ (q : Submodule R M) (e : M ≃L[R] (p × q)), + (∀ x : p, e x = (x, 0)) ∧ (∀ y : q, e y = (0, y)) ∧ (∀ x, e.symm x = x.1 + x.2) := + let ⟨f, hf⟩ := hp + ⟨LinearMap.ker f, .equivOfRightInverse _ p.subtypeL hf, + fun _ ↦ by ext <;> simp [hf], fun _ ↦ by ext <;> simp [hf], fun _ ↦ rfl⟩ + +end Submodule + +namespace MulOpposite + +variable (R : Type*) [Semiring R] [τR : TopologicalSpace R] [TopologicalSemiring R] + {M : Type*} [AddCommMonoid M] [Module R M] [TopologicalSpace M] [ContinuousSMul R M] + +/-- The function `op` is a continuous linear equivalence. -/ +@[simps!] +def opContinuousLinearEquiv : M ≃L[R] Mᵐᵒᵖ where + __ := MulOpposite.opLinearEquiv R + +end MulOpposite + +end diff --git a/Mathlib/Topology/Algebra/Module/LinearMap.lean b/Mathlib/Topology/Algebra/Module/LinearMap.lean new file mode 100644 index 0000000000000..f8fdd3179bce8 --- /dev/null +++ b/Mathlib/Topology/Algebra/Module/LinearMap.lean @@ -0,0 +1,1080 @@ +/- +Copyright (c) 2019 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jan-David Salchow, Sébastien Gouëzel, Jean Lo, Yury Kudryashov, Frédéric Dupuis, + Heather Macbeth +-/ +import Mathlib.Topology.Algebra.Module.Basic +import Mathlib.LinearAlgebra.Projection +import Mathlib.Topology.Algebra.UniformGroup.Defs + +/-! +# Continuous linear maps + +In this file we define continuous (semi-)linear maps, as semilinear maps between topological +modules which are continuous. The set of continuous semilinear maps between the topological +`R₁`-module `M` and `R₂`-module `M₂` with respect to the `RingHom` `σ` is denoted by `M →SL[σ] M₂`. +Plain linear maps are denoted by `M →L[R] M₂` and star-linear maps by `M →L⋆[R] M₂`. +-/ + +assert_not_exists Star.star + +open LinearMap (ker range) +open Topology Filter Pointwise + +universe u v w u' + +/-- Continuous linear maps between modules. We only put the type classes that are necessary for the +definition, although in applications `M` and `M₂` will be topological modules over the topological +ring `R`. -/ +structure ContinuousLinearMap {R : Type*} {S : Type*} [Semiring R] [Semiring S] (σ : R →+* S) + (M : Type*) [TopologicalSpace M] [AddCommMonoid M] (M₂ : Type*) [TopologicalSpace M₂] + [AddCommMonoid M₂] [Module R M] [Module S M₂] extends M →ₛₗ[σ] M₂ where + cont : Continuous toFun := by continuity + +attribute [inherit_doc ContinuousLinearMap] ContinuousLinearMap.cont + +@[inherit_doc] +notation:25 M " →SL[" σ "] " M₂ => ContinuousLinearMap σ M M₂ + +@[inherit_doc] +notation:25 M " →L[" R "] " M₂ => ContinuousLinearMap (RingHom.id R) M M₂ + +/-- `ContinuousSemilinearMapClass F σ M M₂` asserts `F` is a type of bundled continuous +`σ`-semilinear maps `M → M₂`. See also `ContinuousLinearMapClass F R M M₂` for the case where +`σ` is the identity map on `R`. A map `f` between an `R`-module and an `S`-module over a ring +homomorphism `σ : R →+* S` is semilinear if it satisfies the two properties `f (x + y) = f x + f y` +and `f (c • x) = (σ c) • f x`. -/ +class ContinuousSemilinearMapClass (F : Type*) {R S : outParam Type*} [Semiring R] [Semiring S] + (σ : outParam <| R →+* S) (M : outParam Type*) [TopologicalSpace M] [AddCommMonoid M] + (M₂ : outParam Type*) [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R M] + [Module S M₂] [FunLike F M M₂] + extends SemilinearMapClass F σ M M₂, ContinuousMapClass F M M₂ : Prop + +/-- `ContinuousLinearMapClass F R M M₂` asserts `F` is a type of bundled continuous +`R`-linear maps `M → M₂`. This is an abbreviation for +`ContinuousSemilinearMapClass F (RingHom.id R) M M₂`. -/ +abbrev ContinuousLinearMapClass (F : Type*) (R : outParam Type*) [Semiring R] + (M : outParam Type*) [TopologicalSpace M] [AddCommMonoid M] (M₂ : outParam Type*) + [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R M] [Module R M₂] [FunLike F M M₂] := + ContinuousSemilinearMapClass F (RingHom.id R) M M₂ + +namespace ContinuousLinearMap + +section Semiring + +/-! +### Properties that hold for non-necessarily commutative semirings. +-/ + +variable {R₁ : Type*} {R₂ : Type*} {R₃ : Type*} [Semiring R₁] [Semiring R₂] [Semiring R₃] + {σ₁₂ : R₁ →+* R₂} {σ₂₃ : R₂ →+* R₃} {σ₁₃ : R₁ →+* R₃} {M₁ : Type*} [TopologicalSpace M₁] + [AddCommMonoid M₁] {M'₁ : Type*} [TopologicalSpace M'₁] [AddCommMonoid M'₁] {M₂ : Type*} + [TopologicalSpace M₂] [AddCommMonoid M₂] {M₃ : Type*} [TopologicalSpace M₃] [AddCommMonoid M₃] + {M₄ : Type*} [TopologicalSpace M₄] [AddCommMonoid M₄] [Module R₁ M₁] [Module R₁ M'₁] + [Module R₂ M₂] [Module R₃ M₃] + +attribute [coe] ContinuousLinearMap.toLinearMap +/-- Coerce continuous linear maps to linear maps. -/ +instance LinearMap.coe : Coe (M₁ →SL[σ₁₂] M₂) (M₁ →ₛₗ[σ₁₂] M₂) := ⟨toLinearMap⟩ + +theorem coe_injective : Function.Injective ((↑) : (M₁ →SL[σ₁₂] M₂) → M₁ →ₛₗ[σ₁₂] M₂) := by + intro f g H + cases f + cases g + congr + +instance funLike : FunLike (M₁ →SL[σ₁₂] M₂) M₁ M₂ where + coe f := f.toLinearMap + coe_injective' _ _ h := coe_injective (DFunLike.coe_injective h) + +instance continuousSemilinearMapClass : + ContinuousSemilinearMapClass (M₁ →SL[σ₁₂] M₂) σ₁₂ M₁ M₂ where + map_add f := map_add f.toLinearMap + map_continuous f := f.2 + map_smulₛₗ f := f.toLinearMap.map_smul' + +theorem coe_mk (f : M₁ →ₛₗ[σ₁₂] M₂) (h) : (mk f h : M₁ →ₛₗ[σ₁₂] M₂) = f := + rfl + +@[simp] +theorem coe_mk' (f : M₁ →ₛₗ[σ₁₂] M₂) (h) : (mk f h : M₁ → M₂) = f := + rfl + +@[continuity, fun_prop] +protected theorem continuous (f : M₁ →SL[σ₁₂] M₂) : Continuous f := + f.2 + +protected theorem uniformContinuous {E₁ E₂ : Type*} [UniformSpace E₁] [UniformSpace E₂] + [AddCommGroup E₁] [AddCommGroup E₂] [Module R₁ E₁] [Module R₂ E₂] [UniformAddGroup E₁] + [UniformAddGroup E₂] (f : E₁ →SL[σ₁₂] E₂) : UniformContinuous f := + uniformContinuous_addMonoidHom_of_continuous f.continuous + +@[simp, norm_cast] +theorem coe_inj {f g : M₁ →SL[σ₁₂] M₂} : (f : M₁ →ₛₗ[σ₁₂] M₂) = g ↔ f = g := + coe_injective.eq_iff + +theorem coeFn_injective : @Function.Injective (M₁ →SL[σ₁₂] M₂) (M₁ → M₂) (↑) := + DFunLike.coe_injective + +/-- See Note [custom simps projection]. We need to specify this projection explicitly in this case, + because it is a composition of multiple projections. -/ +def Simps.apply (h : M₁ →SL[σ₁₂] M₂) : M₁ → M₂ := + h + +/-- See Note [custom simps projection]. -/ +def Simps.coe (h : M₁ →SL[σ₁₂] M₂) : M₁ →ₛₗ[σ₁₂] M₂ := + h + +initialize_simps_projections ContinuousLinearMap (toFun → apply, toLinearMap → coe) + +@[ext] +theorem ext {f g : M₁ →SL[σ₁₂] M₂} (h : ∀ x, f x = g x) : f = g := + DFunLike.ext f g h + +/-- Copy of a `ContinuousLinearMap` with a new `toFun` equal to the old one. Useful to fix +definitional equalities. -/ +protected def copy (f : M₁ →SL[σ₁₂] M₂) (f' : M₁ → M₂) (h : f' = ⇑f) : M₁ →SL[σ₁₂] M₂ where + toLinearMap := f.toLinearMap.copy f' h + cont := show Continuous f' from h.symm ▸ f.continuous + +@[simp] +theorem coe_copy (f : M₁ →SL[σ₁₂] M₂) (f' : M₁ → M₂) (h : f' = ⇑f) : ⇑(f.copy f' h) = f' := + rfl + +theorem copy_eq (f : M₁ →SL[σ₁₂] M₂) (f' : M₁ → M₂) (h : f' = ⇑f) : f.copy f' h = f := + DFunLike.ext' h + +theorem range_coeFn_eq : + Set.range ((⇑) : (M₁ →SL[σ₁₂] M₂) → (M₁ → M₂)) = + {f | Continuous f} ∩ Set.range ((⇑) : (M₁ →ₛₗ[σ₁₂] M₂) → (M₁ → M₂)) := by + ext f + constructor + · rintro ⟨f, rfl⟩ + exact ⟨f.continuous, f, rfl⟩ + · rintro ⟨hfc, f, rfl⟩ + exact ⟨⟨f, hfc⟩, rfl⟩ + +-- make some straightforward lemmas available to `simp`. +protected theorem map_zero (f : M₁ →SL[σ₁₂] M₂) : f (0 : M₁) = 0 := + map_zero f + +protected theorem map_add (f : M₁ →SL[σ₁₂] M₂) (x y : M₁) : f (x + y) = f x + f y := + map_add f x y + +@[simp] +protected theorem map_smulₛₗ (f : M₁ →SL[σ₁₂] M₂) (c : R₁) (x : M₁) : f (c • x) = σ₁₂ c • f x := + (toLinearMap _).map_smulₛₗ _ _ + +protected theorem map_smul [Module R₁ M₂] (f : M₁ →L[R₁] M₂) (c : R₁) (x : M₁) : + f (c • x) = c • f x := by simp only [RingHom.id_apply, ContinuousLinearMap.map_smulₛₗ] + +@[simp] +theorem map_smul_of_tower {R S : Type*} [Semiring S] [SMul R M₁] [Module S M₁] [SMul R M₂] + [Module S M₂] [LinearMap.CompatibleSMul M₁ M₂ R S] (f : M₁ →L[S] M₂) (c : R) (x : M₁) : + f (c • x) = c • f x := + LinearMap.CompatibleSMul.map_smul (f : M₁ →ₗ[S] M₂) c x + +@[simp, norm_cast] +theorem coe_coe (f : M₁ →SL[σ₁₂] M₂) : ⇑(f : M₁ →ₛₗ[σ₁₂] M₂) = f := + rfl + +@[ext] +theorem ext_ring [TopologicalSpace R₁] {f g : R₁ →L[R₁] M₁} (h : f 1 = g 1) : f = g := + coe_inj.1 <| LinearMap.ext_ring h + +/-- If two continuous linear maps are equal on a set `s`, then they are equal on the closure +of the `Submodule.span` of this set. -/ +theorem eqOn_closure_span [T2Space M₂] {s : Set M₁} {f g : M₁ →SL[σ₁₂] M₂} (h : Set.EqOn f g s) : + Set.EqOn f g (closure (Submodule.span R₁ s : Set M₁)) := + (LinearMap.eqOn_span' h).closure f.continuous g.continuous + +/-- If the submodule generated by a set `s` is dense in the ambient module, then two continuous +linear maps equal on `s` are equal. -/ +theorem ext_on [T2Space M₂] {s : Set M₁} (hs : Dense (Submodule.span R₁ s : Set M₁)) + {f g : M₁ →SL[σ₁₂] M₂} (h : Set.EqOn f g s) : f = g := + ext fun x => eqOn_closure_span h (hs x) + +/-- Under a continuous linear map, the image of the `TopologicalClosure` of a submodule is +contained in the `TopologicalClosure` of its image. -/ +theorem _root_.Submodule.topologicalClosure_map [RingHomSurjective σ₁₂] [TopologicalSpace R₁] + [TopologicalSpace R₂] [ContinuousSMul R₁ M₁] [ContinuousAdd M₁] [ContinuousSMul R₂ M₂] + [ContinuousAdd M₂] (f : M₁ →SL[σ₁₂] M₂) (s : Submodule R₁ M₁) : + s.topologicalClosure.map (f : M₁ →ₛₗ[σ₁₂] M₂) ≤ + (s.map (f : M₁ →ₛₗ[σ₁₂] M₂)).topologicalClosure := + image_closure_subset_closure_image f.continuous + +/-- Under a dense continuous linear map, a submodule whose `TopologicalClosure` is `⊤` is sent to +another such submodule. That is, the image of a dense set under a map with dense range is dense. +-/ +theorem _root_.DenseRange.topologicalClosure_map_submodule [RingHomSurjective σ₁₂] + [TopologicalSpace R₁] [TopologicalSpace R₂] [ContinuousSMul R₁ M₁] [ContinuousAdd M₁] + [ContinuousSMul R₂ M₂] [ContinuousAdd M₂] {f : M₁ →SL[σ₁₂] M₂} (hf' : DenseRange f) + {s : Submodule R₁ M₁} (hs : s.topologicalClosure = ⊤) : + (s.map (f : M₁ →ₛₗ[σ₁₂] M₂)).topologicalClosure = ⊤ := by + rw [SetLike.ext'_iff] at hs ⊢ + simp only [Submodule.topologicalClosure_coe, Submodule.top_coe, ← dense_iff_closure_eq] at hs ⊢ + exact hf'.dense_image f.continuous hs + +section SMulMonoid + +variable {S₂ T₂ : Type*} [Monoid S₂] [Monoid T₂] +variable [DistribMulAction S₂ M₂] [SMulCommClass R₂ S₂ M₂] [ContinuousConstSMul S₂ M₂] +variable [DistribMulAction T₂ M₂] [SMulCommClass R₂ T₂ M₂] [ContinuousConstSMul T₂ M₂] + +instance instSMul : SMul S₂ (M₁ →SL[σ₁₂] M₂) where + smul c f := ⟨c • (f : M₁ →ₛₗ[σ₁₂] M₂), (f.2.const_smul _ : Continuous fun x => c • f x)⟩ + +instance mulAction : MulAction S₂ (M₁ →SL[σ₁₂] M₂) where + one_smul _f := ext fun _x => one_smul _ _ + mul_smul _a _b _f := ext fun _x => mul_smul _ _ _ + +theorem smul_apply (c : S₂) (f : M₁ →SL[σ₁₂] M₂) (x : M₁) : (c • f) x = c • f x := + rfl + +@[simp, norm_cast] +theorem coe_smul (c : S₂) (f : M₁ →SL[σ₁₂] M₂) : + ↑(c • f) = c • (f : M₁ →ₛₗ[σ₁₂] M₂) := + rfl + +@[simp, norm_cast] +theorem coe_smul' (c : S₂) (f : M₁ →SL[σ₁₂] M₂) : + ↑(c • f) = c • (f : M₁ → M₂) := + rfl + +instance isScalarTower [SMul S₂ T₂] [IsScalarTower S₂ T₂ M₂] : + IsScalarTower S₂ T₂ (M₁ →SL[σ₁₂] M₂) := + ⟨fun a b f => ext fun x => smul_assoc a b (f x)⟩ + +instance smulCommClass [SMulCommClass S₂ T₂ M₂] : SMulCommClass S₂ T₂ (M₁ →SL[σ₁₂] M₂) := + ⟨fun a b f => ext fun x => smul_comm a b (f x)⟩ + +end SMulMonoid + +/-- The continuous map that is constantly zero. -/ +instance zero : Zero (M₁ →SL[σ₁₂] M₂) := + ⟨⟨0, continuous_zero⟩⟩ + +instance inhabited : Inhabited (M₁ →SL[σ₁₂] M₂) := + ⟨0⟩ + +@[simp] +theorem default_def : (default : M₁ →SL[σ₁₂] M₂) = 0 := + rfl + +@[simp] +theorem zero_apply (x : M₁) : (0 : M₁ →SL[σ₁₂] M₂) x = 0 := + rfl + +@[simp, norm_cast] +theorem coe_zero : ((0 : M₁ →SL[σ₁₂] M₂) : M₁ →ₛₗ[σ₁₂] M₂) = 0 := + rfl + +/- no simp attribute on the next line as simp does not always simplify `0 x` to `0` +when `0` is the zero function, while it does for the zero continuous linear map, +and this is the most important property we care about. -/ +@[norm_cast] +theorem coe_zero' : ⇑(0 : M₁ →SL[σ₁₂] M₂) = 0 := + rfl + +instance uniqueOfLeft [Subsingleton M₁] : Unique (M₁ →SL[σ₁₂] M₂) := + coe_injective.unique + +instance uniqueOfRight [Subsingleton M₂] : Unique (M₁ →SL[σ₁₂] M₂) := + coe_injective.unique + +theorem exists_ne_zero {f : M₁ →SL[σ₁₂] M₂} (hf : f ≠ 0) : ∃ x, f x ≠ 0 := by + by_contra! h + exact hf (ContinuousLinearMap.ext h) + +section + +variable (R₁ M₁) + +/-- the identity map as a continuous linear map. -/ +def id : M₁ →L[R₁] M₁ := + ⟨LinearMap.id, continuous_id⟩ + +end + +instance one : One (M₁ →L[R₁] M₁) := + ⟨id R₁ M₁⟩ + +theorem one_def : (1 : M₁ →L[R₁] M₁) = id R₁ M₁ := + rfl + +theorem id_apply (x : M₁) : id R₁ M₁ x = x := + rfl + +@[simp, norm_cast] +theorem coe_id : (id R₁ M₁ : M₁ →ₗ[R₁] M₁) = LinearMap.id := + rfl + +@[simp, norm_cast] +theorem coe_id' : ⇑(id R₁ M₁) = _root_.id := + rfl + +@[simp, norm_cast] +theorem coe_eq_id {f : M₁ →L[R₁] M₁} : (f : M₁ →ₗ[R₁] M₁) = LinearMap.id ↔ f = id _ _ := by + rw [← coe_id, coe_inj] + +@[simp] +theorem one_apply (x : M₁) : (1 : M₁ →L[R₁] M₁) x = x := + rfl + +instance [Nontrivial M₁] : Nontrivial (M₁ →L[R₁] M₁) := + ⟨0, 1, fun e ↦ + have ⟨x, hx⟩ := exists_ne (0 : M₁); hx (by simpa using DFunLike.congr_fun e.symm x)⟩ + +section Add + +variable [ContinuousAdd M₂] + +instance add : Add (M₁ →SL[σ₁₂] M₂) := + ⟨fun f g => ⟨f + g, f.2.add g.2⟩⟩ + +@[simp] +theorem add_apply (f g : M₁ →SL[σ₁₂] M₂) (x : M₁) : (f + g) x = f x + g x := + rfl + +@[simp, norm_cast] +theorem coe_add (f g : M₁ →SL[σ₁₂] M₂) : (↑(f + g) : M₁ →ₛₗ[σ₁₂] M₂) = f + g := + rfl + +@[norm_cast] +theorem coe_add' (f g : M₁ →SL[σ₁₂] M₂) : ⇑(f + g) = f + g := + rfl + +instance addCommMonoid : AddCommMonoid (M₁ →SL[σ₁₂] M₂) where + zero_add := by + intros + ext + apply_rules [zero_add, add_assoc, add_zero, neg_add_cancel, add_comm] + add_zero := by + intros + ext + apply_rules [zero_add, add_assoc, add_zero, neg_add_cancel, add_comm] + add_comm := by + intros + ext + apply_rules [zero_add, add_assoc, add_zero, neg_add_cancel, add_comm] + add_assoc := by + intros + ext + apply_rules [zero_add, add_assoc, add_zero, neg_add_cancel, add_comm] + nsmul := (· • ·) + nsmul_zero f := by + ext + simp + nsmul_succ n f := by + ext + simp [add_smul] + +@[simp, norm_cast] +theorem coe_sum {ι : Type*} (t : Finset ι) (f : ι → M₁ →SL[σ₁₂] M₂) : + ↑(∑ d ∈ t, f d) = (∑ d ∈ t, f d : M₁ →ₛₗ[σ₁₂] M₂) := + map_sum (AddMonoidHom.mk ⟨((↑) : (M₁ →SL[σ₁₂] M₂) → M₁ →ₛₗ[σ₁₂] M₂), rfl⟩ fun _ _ => rfl) _ _ + +@[simp, norm_cast] +theorem coe_sum' {ι : Type*} (t : Finset ι) (f : ι → M₁ →SL[σ₁₂] M₂) : + ⇑(∑ d ∈ t, f d) = ∑ d ∈ t, ⇑(f d) := by simp only [← coe_coe, coe_sum, LinearMap.coeFn_sum] + +theorem sum_apply {ι : Type*} (t : Finset ι) (f : ι → M₁ →SL[σ₁₂] M₂) (b : M₁) : + (∑ d ∈ t, f d) b = ∑ d ∈ t, f d b := by simp only [coe_sum', Finset.sum_apply] + +end Add + +variable [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] + +/-- Composition of bounded linear maps. -/ +def comp (g : M₂ →SL[σ₂₃] M₃) (f : M₁ →SL[σ₁₂] M₂) : M₁ →SL[σ₁₃] M₃ := + ⟨(g : M₂ →ₛₗ[σ₂₃] M₃).comp (f : M₁ →ₛₗ[σ₁₂] M₂), g.2.comp f.2⟩ + +@[inherit_doc comp] +infixr:80 " ∘L " => + @ContinuousLinearMap.comp _ _ _ _ _ _ (RingHom.id _) (RingHom.id _) (RingHom.id _) _ _ _ _ _ _ _ _ + _ _ _ _ RingHomCompTriple.ids + +@[simp, norm_cast] +theorem coe_comp (h : M₂ →SL[σ₂₃] M₃) (f : M₁ →SL[σ₁₂] M₂) : + (h.comp f : M₁ →ₛₗ[σ₁₃] M₃) = (h : M₂ →ₛₗ[σ₂₃] M₃).comp (f : M₁ →ₛₗ[σ₁₂] M₂) := + rfl + +@[simp, norm_cast] +theorem coe_comp' (h : M₂ →SL[σ₂₃] M₃) (f : M₁ →SL[σ₁₂] M₂) : ⇑(h.comp f) = h ∘ f := + rfl + +theorem comp_apply (g : M₂ →SL[σ₂₃] M₃) (f : M₁ →SL[σ₁₂] M₂) (x : M₁) : (g.comp f) x = g (f x) := + rfl + +@[simp] +theorem comp_id (f : M₁ →SL[σ₁₂] M₂) : f.comp (id R₁ M₁) = f := + ext fun _x => rfl + +@[simp] +theorem id_comp (f : M₁ →SL[σ₁₂] M₂) : (id R₂ M₂).comp f = f := + ext fun _x => rfl + +@[simp] +theorem comp_zero (g : M₂ →SL[σ₂₃] M₃) : g.comp (0 : M₁ →SL[σ₁₂] M₂) = 0 := by + ext + simp + +@[simp] +theorem zero_comp (f : M₁ →SL[σ₁₂] M₂) : (0 : M₂ →SL[σ₂₃] M₃).comp f = 0 := by + ext + simp + +@[simp] +theorem comp_add [ContinuousAdd M₂] [ContinuousAdd M₃] (g : M₂ →SL[σ₂₃] M₃) + (f₁ f₂ : M₁ →SL[σ₁₂] M₂) : g.comp (f₁ + f₂) = g.comp f₁ + g.comp f₂ := by + ext + simp + +@[simp] +theorem add_comp [ContinuousAdd M₃] (g₁ g₂ : M₂ →SL[σ₂₃] M₃) (f : M₁ →SL[σ₁₂] M₂) : + (g₁ + g₂).comp f = g₁.comp f + g₂.comp f := by + ext + simp + +theorem comp_finset_sum {ι : Type*} {s : Finset ι} + [ContinuousAdd M₂] [ContinuousAdd M₃] (g : M₂ →SL[σ₂₃] M₃) + (f : ι → M₁ →SL[σ₁₂] M₂) : g.comp (∑ i ∈ s, f i) = ∑ i ∈ s, g.comp (f i) := by + ext + simp + +theorem finset_sum_comp {ι : Type*} {s : Finset ι} + [ContinuousAdd M₃] (g : ι → M₂ →SL[σ₂₃] M₃) + (f : M₁ →SL[σ₁₂] M₂) : (∑ i ∈ s, g i).comp f = ∑ i ∈ s, (g i).comp f := by + ext + simp only [coe_comp', coe_sum', Function.comp_apply, Finset.sum_apply] + +theorem comp_assoc {R₄ : Type*} [Semiring R₄] [Module R₄ M₄] {σ₁₄ : R₁ →+* R₄} {σ₂₄ : R₂ →+* R₄} + {σ₃₄ : R₃ →+* R₄} [RingHomCompTriple σ₁₃ σ₃₄ σ₁₄] [RingHomCompTriple σ₂₃ σ₃₄ σ₂₄] + [RingHomCompTriple σ₁₂ σ₂₄ σ₁₄] (h : M₃ →SL[σ₃₄] M₄) (g : M₂ →SL[σ₂₃] M₃) (f : M₁ →SL[σ₁₂] M₂) : + (h.comp g).comp f = h.comp (g.comp f) := + rfl + +instance instMul : Mul (M₁ →L[R₁] M₁) := + ⟨comp⟩ + +theorem mul_def (f g : M₁ →L[R₁] M₁) : f * g = f.comp g := + rfl + +@[simp] +theorem coe_mul (f g : M₁ →L[R₁] M₁) : ⇑(f * g) = f ∘ g := + rfl + +theorem mul_apply (f g : M₁ →L[R₁] M₁) (x : M₁) : (f * g) x = f (g x) := + rfl + +instance monoidWithZero : MonoidWithZero (M₁ →L[R₁] M₁) where + mul_zero f := ext fun _ => map_zero f + zero_mul _ := ext fun _ => rfl + mul_one _ := ext fun _ => rfl + one_mul _ := ext fun _ => rfl + mul_assoc _ _ _ := ext fun _ => rfl + +theorem coe_pow (f : M₁ →L[R₁] M₁) (n : ℕ) : ⇑(f ^ n) = f^[n] := + hom_coe_pow _ rfl (fun _ _ ↦ rfl) _ _ + +instance instNatCast [ContinuousAdd M₁] : NatCast (M₁ →L[R₁] M₁) where + natCast n := n • (1 : M₁ →L[R₁] M₁) + +instance semiring [ContinuousAdd M₁] : Semiring (M₁ →L[R₁] M₁) where + __ := ContinuousLinearMap.monoidWithZero + __ := ContinuousLinearMap.addCommMonoid + left_distrib f g h := ext fun x => map_add f (g x) (h x) + right_distrib _ _ _ := ext fun _ => LinearMap.add_apply _ _ _ + toNatCast := instNatCast + natCast_zero := zero_smul ℕ (1 : M₁ →L[R₁] M₁) + natCast_succ n := AddMonoid.nsmul_succ n (1 : M₁ →L[R₁] M₁) + +/-- `ContinuousLinearMap.toLinearMap` as a `RingHom`. -/ +@[simps] +def toLinearMapRingHom [ContinuousAdd M₁] : (M₁ →L[R₁] M₁) →+* M₁ →ₗ[R₁] M₁ where + toFun := toLinearMap + map_zero' := rfl + map_one' := rfl + map_add' _ _ := rfl + map_mul' _ _ := rfl + +@[simp] +theorem natCast_apply [ContinuousAdd M₁] (n : ℕ) (m : M₁) : (↑n : M₁ →L[R₁] M₁) m = n • m := + rfl + +@[simp] +theorem ofNat_apply [ContinuousAdd M₁] (n : ℕ) [n.AtLeastTwo] (m : M₁) : + ((no_index (OfNat.ofNat n) : M₁ →L[R₁] M₁)) m = OfNat.ofNat n • m := + rfl + +section ApplyAction + +variable [ContinuousAdd M₁] + +/-- The tautological action by `M₁ →L[R₁] M₁` on `M`. + +This generalizes `Function.End.applyMulAction`. -/ +instance applyModule : Module (M₁ →L[R₁] M₁) M₁ := + Module.compHom _ toLinearMapRingHom + +@[simp] +protected theorem smul_def (f : M₁ →L[R₁] M₁) (a : M₁) : f • a = f a := + rfl + +/-- `ContinuousLinearMap.applyModule` is faithful. -/ +instance applyFaithfulSMul : FaithfulSMul (M₁ →L[R₁] M₁) M₁ := + ⟨fun {_ _} => ContinuousLinearMap.ext⟩ + +instance applySMulCommClass : SMulCommClass R₁ (M₁ →L[R₁] M₁) M₁ where + smul_comm r e m := (e.map_smul r m).symm + +instance applySMulCommClass' : SMulCommClass (M₁ →L[R₁] M₁) R₁ M₁ where + smul_comm := ContinuousLinearMap.map_smul + +instance continuousConstSMul_apply : ContinuousConstSMul (M₁ →L[R₁] M₁) M₁ := + ⟨ContinuousLinearMap.continuous⟩ + +end ApplyAction + +variable {F : Type*} + +theorem isClosed_ker [T1Space M₂] [FunLike F M₁ M₂] [ContinuousSemilinearMapClass F σ₁₂ M₁ M₂] + (f : F) : + IsClosed (ker f : Set M₁) := + continuous_iff_isClosed.1 (map_continuous f) _ isClosed_singleton + +theorem isComplete_ker {M' : Type*} [UniformSpace M'] [CompleteSpace M'] [AddCommMonoid M'] + [Module R₁ M'] [T1Space M₂] [FunLike F M' M₂] [ContinuousSemilinearMapClass F σ₁₂ M' M₂] + (f : F) : + IsComplete (ker f : Set M') := + (isClosed_ker f).isComplete + +instance completeSpace_ker {M' : Type*} [UniformSpace M'] [CompleteSpace M'] + [AddCommMonoid M'] [Module R₁ M'] [T1Space M₂] + [FunLike F M' M₂] [ContinuousSemilinearMapClass F σ₁₂ M' M₂] + (f : F) : CompleteSpace (ker f) := + (isComplete_ker f).completeSpace_coe + +instance completeSpace_eqLocus {M' : Type*} [UniformSpace M'] [CompleteSpace M'] + [AddCommMonoid M'] [Module R₁ M'] [T2Space M₂] + [FunLike F M' M₂] [ContinuousSemilinearMapClass F σ₁₂ M' M₂] + (f g : F) : CompleteSpace (LinearMap.eqLocus f g) := + IsClosed.completeSpace_coe <| isClosed_eq (map_continuous f) (map_continuous g) + +/-- Restrict codomain of a continuous linear map. -/ +def codRestrict (f : M₁ →SL[σ₁₂] M₂) (p : Submodule R₂ M₂) (h : ∀ x, f x ∈ p) : + M₁ →SL[σ₁₂] p where + cont := f.continuous.subtype_mk _ + toLinearMap := (f : M₁ →ₛₗ[σ₁₂] M₂).codRestrict p h + +@[norm_cast] +theorem coe_codRestrict (f : M₁ →SL[σ₁₂] M₂) (p : Submodule R₂ M₂) (h : ∀ x, f x ∈ p) : + (f.codRestrict p h : M₁ →ₛₗ[σ₁₂] p) = (f : M₁ →ₛₗ[σ₁₂] M₂).codRestrict p h := + rfl + +@[simp] +theorem coe_codRestrict_apply (f : M₁ →SL[σ₁₂] M₂) (p : Submodule R₂ M₂) (h : ∀ x, f x ∈ p) (x) : + (f.codRestrict p h x : M₂) = f x := + rfl + +@[simp] +theorem ker_codRestrict (f : M₁ →SL[σ₁₂] M₂) (p : Submodule R₂ M₂) (h : ∀ x, f x ∈ p) : + ker (f.codRestrict p h) = ker f := + (f : M₁ →ₛₗ[σ₁₂] M₂).ker_codRestrict p h + +/-- Restrict the codomain of a continuous linear map `f` to `f.range`. -/ +abbrev rangeRestrict [RingHomSurjective σ₁₂] (f : M₁ →SL[σ₁₂] M₂) := + f.codRestrict (LinearMap.range f) (LinearMap.mem_range_self f) + +@[simp] +theorem coe_rangeRestrict [RingHomSurjective σ₁₂] (f : M₁ →SL[σ₁₂] M₂) : + (f.rangeRestrict : M₁ →ₛₗ[σ₁₂] LinearMap.range f) = (f : M₁ →ₛₗ[σ₁₂] M₂).rangeRestrict := + rfl + +/-- `Submodule.subtype` as a `ContinuousLinearMap`. -/ +def _root_.Submodule.subtypeL (p : Submodule R₁ M₁) : p →L[R₁] M₁ where + cont := continuous_subtype_val + toLinearMap := p.subtype + +@[simp, norm_cast] +theorem _root_.Submodule.coe_subtypeL (p : Submodule R₁ M₁) : + (p.subtypeL : p →ₗ[R₁] M₁) = p.subtype := + rfl + +@[simp] +theorem _root_.Submodule.coe_subtypeL' (p : Submodule R₁ M₁) : ⇑p.subtypeL = p.subtype := + rfl + +@[simp] -- @[norm_cast] -- Porting note: A theorem with this can't have a rhs starting with `↑`. +theorem _root_.Submodule.subtypeL_apply (p : Submodule R₁ M₁) (x : p) : p.subtypeL x = x := + rfl + +@[simp] +theorem _root_.Submodule.range_subtypeL (p : Submodule R₁ M₁) : range p.subtypeL = p := + Submodule.range_subtype _ + +@[simp] +theorem _root_.Submodule.ker_subtypeL (p : Submodule R₁ M₁) : ker p.subtypeL = ⊥ := + Submodule.ker_subtype _ + +section + +variable {R S : Type*} [Semiring R] [Semiring S] [Module R M₁] [Module R M₂] [Module R S] + [Module S M₂] [IsScalarTower R S M₂] [TopologicalSpace S] [ContinuousSMul S M₂] + +/-- The linear map `fun x => c x • f`. Associates to a scalar-valued linear map and an element of +`M₂` the `M₂`-valued linear map obtained by multiplying the two (a.k.a. tensoring by `M₂`). +See also `ContinuousLinearMap.smulRightₗ` and `ContinuousLinearMap.smulRightL`. -/ +def smulRight (c : M₁ →L[R] S) (f : M₂) : M₁ →L[R] M₂ := + { c.toLinearMap.smulRight f with cont := c.2.smul continuous_const } + +@[simp] +theorem smulRight_apply {c : M₁ →L[R] S} {f : M₂} {x : M₁} : + (smulRight c f : M₁ → M₂) x = c x • f := + rfl + +end + +variable [Module R₁ M₂] [TopologicalSpace R₁] [ContinuousSMul R₁ M₂] + +@[simp] +theorem smulRight_one_one (c : R₁ →L[R₁] M₂) : smulRight (1 : R₁ →L[R₁] R₁) (c 1) = c := by + ext + simp [← ContinuousLinearMap.map_smul_of_tower] + +@[simp] +theorem smulRight_one_eq_iff {f f' : M₂} : + smulRight (1 : R₁ →L[R₁] R₁) f = smulRight (1 : R₁ →L[R₁] R₁) f' ↔ f = f' := by + simp only [ContinuousLinearMap.ext_ring_iff, smulRight_apply, one_apply, one_smul] + +theorem smulRight_comp [ContinuousMul R₁] {x : M₂} {c : R₁} : + (smulRight (1 : R₁ →L[R₁] R₁) x).comp (smulRight (1 : R₁ →L[R₁] R₁) c) = + smulRight (1 : R₁ →L[R₁] R₁) (c • x) := by + ext + simp [mul_smul] + +section ToSpanSingleton + +variable (R₁) +variable [ContinuousSMul R₁ M₁] + +/-- Given an element `x` of a topological space `M` over a semiring `R`, the natural continuous +linear map from `R` to `M` by taking multiples of `x`. -/ +def toSpanSingleton (x : M₁) : R₁ →L[R₁] M₁ where + toLinearMap := LinearMap.toSpanSingleton R₁ M₁ x + cont := continuous_id.smul continuous_const + +theorem toSpanSingleton_apply (x : M₁) (r : R₁) : toSpanSingleton R₁ x r = r • x := + rfl + +theorem toSpanSingleton_add [ContinuousAdd M₁] (x y : M₁) : + toSpanSingleton R₁ (x + y) = toSpanSingleton R₁ x + toSpanSingleton R₁ y := by + ext1; simp [toSpanSingleton_apply] + +theorem toSpanSingleton_smul' {α} [Monoid α] [DistribMulAction α M₁] [ContinuousConstSMul α M₁] + [SMulCommClass R₁ α M₁] (c : α) (x : M₁) : + toSpanSingleton R₁ (c • x) = c • toSpanSingleton R₁ x := by + ext1; rw [toSpanSingleton_apply, smul_apply, toSpanSingleton_apply, smul_comm] + +/-- A special case of `to_span_singleton_smul'` for when `R` is commutative. -/ +theorem toSpanSingleton_smul (R) {M₁} [CommSemiring R] [AddCommMonoid M₁] [Module R M₁] + [TopologicalSpace R] [TopologicalSpace M₁] [ContinuousSMul R M₁] (c : R) (x : M₁) : + toSpanSingleton R (c • x) = c • toSpanSingleton R x := + toSpanSingleton_smul' R c x + +end ToSpanSingleton + +end Semiring + +section Ring + +variable {R : Type*} [Ring R] {R₂ : Type*} [Ring R₂] {R₃ : Type*} [Ring R₃] {M : Type*} + [TopologicalSpace M] [AddCommGroup M] {M₂ : Type*} [TopologicalSpace M₂] [AddCommGroup M₂] + {M₃ : Type*} [TopologicalSpace M₃] [AddCommGroup M₃] {M₄ : Type*} [TopologicalSpace M₄] + [AddCommGroup M₄] [Module R M] [Module R₂ M₂] [Module R₃ M₃] {σ₁₂ : R →+* R₂} {σ₂₃ : R₂ →+* R₃} + {σ₁₃ : R →+* R₃} + +section + +protected theorem map_neg (f : M →SL[σ₁₂] M₂) (x : M) : f (-x) = -f x := by + exact map_neg f x + +protected theorem map_sub (f : M →SL[σ₁₂] M₂) (x y : M) : f (x - y) = f x - f y := by + exact map_sub f x y + +@[simp] +theorem sub_apply' (f g : M →SL[σ₁₂] M₂) (x : M) : ((f : M →ₛₗ[σ₁₂] M₂) - g) x = f x - g x := + rfl + +end + +section + +variable [TopologicalAddGroup M₂] + +instance neg : Neg (M →SL[σ₁₂] M₂) := + ⟨fun f => ⟨-f, f.2.neg⟩⟩ + +@[simp] +theorem neg_apply (f : M →SL[σ₁₂] M₂) (x : M) : (-f) x = -f x := + rfl + +@[simp, norm_cast] +theorem coe_neg (f : M →SL[σ₁₂] M₂) : (↑(-f) : M →ₛₗ[σ₁₂] M₂) = -f := + rfl + +@[norm_cast] +theorem coe_neg' (f : M →SL[σ₁₂] M₂) : ⇑(-f) = -f := + rfl + +instance sub : Sub (M →SL[σ₁₂] M₂) := + ⟨fun f g => ⟨f - g, f.2.sub g.2⟩⟩ + +instance addCommGroup : AddCommGroup (M →SL[σ₁₂] M₂) where + __ := ContinuousLinearMap.addCommMonoid + neg := (-·) + sub := (· - ·) + sub_eq_add_neg _ _ := by ext; apply sub_eq_add_neg + nsmul := (· • ·) + zsmul := (· • ·) + zsmul_zero' f := by ext; simp + zsmul_succ' n f := by ext; simp [add_smul, add_comm] + zsmul_neg' n f := by ext; simp [add_smul] + neg_add_cancel _ := by ext; apply neg_add_cancel + +theorem sub_apply (f g : M →SL[σ₁₂] M₂) (x : M) : (f - g) x = f x - g x := + rfl + +@[simp, norm_cast] +theorem coe_sub (f g : M →SL[σ₁₂] M₂) : (↑(f - g) : M →ₛₗ[σ₁₂] M₂) = f - g := + rfl + +@[simp, norm_cast] +theorem coe_sub' (f g : M →SL[σ₁₂] M₂) : ⇑(f - g) = f - g := + rfl + +end + +@[simp] +theorem comp_neg [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] [TopologicalAddGroup M₂] [TopologicalAddGroup M₃] + (g : M₂ →SL[σ₂₃] M₃) (f : M →SL[σ₁₂] M₂) : g.comp (-f) = -g.comp f := by + ext x + simp + +@[simp] +theorem neg_comp [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] [TopologicalAddGroup M₃] (g : M₂ →SL[σ₂₃] M₃) + (f : M →SL[σ₁₂] M₂) : (-g).comp f = -g.comp f := by + ext + simp + +@[simp] +theorem comp_sub [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] [TopologicalAddGroup M₂] [TopologicalAddGroup M₃] + (g : M₂ →SL[σ₂₃] M₃) (f₁ f₂ : M →SL[σ₁₂] M₂) : g.comp (f₁ - f₂) = g.comp f₁ - g.comp f₂ := by + ext + simp + +@[simp] +theorem sub_comp [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] [TopologicalAddGroup M₃] (g₁ g₂ : M₂ →SL[σ₂₃] M₃) + (f : M →SL[σ₁₂] M₂) : (g₁ - g₂).comp f = g₁.comp f - g₂.comp f := by + ext + simp + +instance ring [TopologicalAddGroup M] : Ring (M →L[R] M) where + __ := ContinuousLinearMap.semiring + __ := ContinuousLinearMap.addCommGroup + intCast z := z • (1 : M →L[R] M) + intCast_ofNat := natCast_zsmul _ + intCast_negSucc := negSucc_zsmul _ + +@[simp] +theorem intCast_apply [TopologicalAddGroup M] (z : ℤ) (m : M) : (↑z : M →L[R] M) m = z • m := + rfl + +theorem smulRight_one_pow [TopologicalSpace R] [TopologicalRing R] (c : R) (n : ℕ) : + smulRight (1 : R →L[R] R) c ^ n = smulRight (1 : R →L[R] R) (c ^ n) := by + induction n with + | zero => ext; simp + | succ n ihn => rw [pow_succ, ihn, mul_def, smulRight_comp, smul_eq_mul, pow_succ'] + +section + +variable {σ₂₁ : R₂ →+* R} [RingHomInvPair σ₁₂ σ₂₁] + + +/-- Given a right inverse `f₂ : M₂ →L[R] M` to `f₁ : M →L[R] M₂`, +`projKerOfRightInverse f₁ f₂ h` is the projection `M →L[R] LinearMap.ker f₁` along +`LinearMap.range f₂`. -/ +def projKerOfRightInverse [TopologicalAddGroup M] (f₁ : M →SL[σ₁₂] M₂) (f₂ : M₂ →SL[σ₂₁] M) + (h : Function.RightInverse f₂ f₁) : M →L[R] LinearMap.ker f₁ := + (id R M - f₂.comp f₁).codRestrict (LinearMap.ker f₁) fun x => by simp [h (f₁ x)] + +@[simp] +theorem coe_projKerOfRightInverse_apply [TopologicalAddGroup M] (f₁ : M →SL[σ₁₂] M₂) + (f₂ : M₂ →SL[σ₂₁] M) (h : Function.RightInverse f₂ f₁) (x : M) : + (f₁.projKerOfRightInverse f₂ h x : M) = x - f₂ (f₁ x) := + rfl + +@[simp] +theorem projKerOfRightInverse_apply_idem [TopologicalAddGroup M] (f₁ : M →SL[σ₁₂] M₂) + (f₂ : M₂ →SL[σ₂₁] M) (h : Function.RightInverse f₂ f₁) (x : LinearMap.ker f₁) : + f₁.projKerOfRightInverse f₂ h x = x := by + ext1 + simp + +@[simp] +theorem projKerOfRightInverse_comp_inv [TopologicalAddGroup M] (f₁ : M →SL[σ₁₂] M₂) + (f₂ : M₂ →SL[σ₂₁] M) (h : Function.RightInverse f₂ f₁) (y : M₂) : + f₁.projKerOfRightInverse f₂ h (f₂ y) = 0 := + Subtype.ext_iff_val.2 <| by simp [h y] + +end + +end Ring + +section DivisionMonoid + +variable {R M : Type*} + +/-- A nonzero continuous linear functional is open. -/ +protected theorem isOpenMap_of_ne_zero [TopologicalSpace R] [DivisionRing R] [ContinuousSub R] + [AddCommGroup M] [TopologicalSpace M] [ContinuousAdd M] [Module R M] [ContinuousSMul R M] + (f : M →L[R] R) (hf : f ≠ 0) : IsOpenMap f := + let ⟨x, hx⟩ := exists_ne_zero hf + IsOpenMap.of_sections fun y => + ⟨fun a => y + (a - f y) • (f x)⁻¹ • x, Continuous.continuousAt <| by continuity, by simp, + fun a => by simp [hx]⟩ + +end DivisionMonoid + +section SMulMonoid + +-- The M's are used for semilinear maps, and the N's for plain linear maps +variable {R R₂ R₃ S S₃ : Type*} [Semiring R] [Semiring R₂] [Semiring R₃] [Monoid S] [Monoid S₃] + {M : Type*} [TopologicalSpace M] [AddCommMonoid M] [Module R M] {M₂ : Type*} + [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₂ M₂] {M₃ : Type*} [TopologicalSpace M₃] + [AddCommMonoid M₃] [Module R₃ M₃] {N₂ : Type*} [TopologicalSpace N₂] [AddCommMonoid N₂] + [Module R N₂] {N₃ : Type*} [TopologicalSpace N₃] [AddCommMonoid N₃] [Module R N₃] + [DistribMulAction S₃ M₃] [SMulCommClass R₃ S₃ M₃] [ContinuousConstSMul S₃ M₃] + [DistribMulAction S N₃] [SMulCommClass R S N₃] [ContinuousConstSMul S N₃] {σ₁₂ : R →+* R₂} + {σ₂₃ : R₂ →+* R₃} {σ₁₃ : R →+* R₃} [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] + +@[simp] +theorem smul_comp (c : S₃) (h : M₂ →SL[σ₂₃] M₃) (f : M →SL[σ₁₂] M₂) : + (c • h).comp f = c • h.comp f := + rfl + +variable [DistribMulAction S₃ M₂] [ContinuousConstSMul S₃ M₂] [SMulCommClass R₂ S₃ M₂] +variable [DistribMulAction S N₂] [ContinuousConstSMul S N₂] [SMulCommClass R S N₂] + +@[simp] +theorem comp_smul [LinearMap.CompatibleSMul N₂ N₃ S R] (hₗ : N₂ →L[R] N₃) (c : S) + (fₗ : M →L[R] N₂) : hₗ.comp (c • fₗ) = c • hₗ.comp fₗ := by + ext x + exact hₗ.map_smul_of_tower c (fₗ x) + +@[simp] +theorem comp_smulₛₗ [SMulCommClass R₂ R₂ M₂] [SMulCommClass R₃ R₃ M₃] [ContinuousConstSMul R₂ M₂] + [ContinuousConstSMul R₃ M₃] (h : M₂ →SL[σ₂₃] M₃) (c : R₂) (f : M →SL[σ₁₂] M₂) : + h.comp (c • f) = σ₂₃ c • h.comp f := by + ext x + simp only [coe_smul', coe_comp', Function.comp_apply, Pi.smul_apply, + ContinuousLinearMap.map_smulₛₗ] + +instance distribMulAction [ContinuousAdd M₂] : DistribMulAction S₃ (M →SL[σ₁₂] M₂) where + smul_add a f g := ext fun x => smul_add a (f x) (g x) + smul_zero a := ext fun _ => smul_zero a + +end SMulMonoid + +section SMul + +-- The M's are used for semilinear maps, and the N's for plain linear maps +variable {R R₂ R₃ S S₃ : Type*} [Semiring R] [Semiring R₂] [Semiring R₃] [Semiring S] [Semiring S₃] + {M : Type*} [TopologicalSpace M] [AddCommMonoid M] [Module R M] {M₂ : Type*} + [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R₂ M₂] {M₃ : Type*} [TopologicalSpace M₃] + [AddCommMonoid M₃] [Module R₃ M₃] {N₂ : Type*} [TopologicalSpace N₂] [AddCommMonoid N₂] + [Module R N₂] {N₃ : Type*} [TopologicalSpace N₃] [AddCommMonoid N₃] [Module R N₃] [Module S₃ M₃] + [SMulCommClass R₃ S₃ M₃] [ContinuousConstSMul S₃ M₃] [Module S N₂] [ContinuousConstSMul S N₂] + [SMulCommClass R S N₂] [Module S N₃] [SMulCommClass R S N₃] [ContinuousConstSMul S N₃] + {σ₁₂ : R →+* R₂} {σ₂₃ : R₂ →+* R₃} {σ₁₃ : R →+* R₃} [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] (c : S) + (h : M₂ →SL[σ₂₃] M₃) (f : M →SL[σ₁₂] M₂) + +variable [ContinuousAdd M₂] [ContinuousAdd M₃] [ContinuousAdd N₂] + +instance module : Module S₃ (M →SL[σ₁₃] M₃) where + zero_smul _ := ext fun _ => zero_smul S₃ _ + add_smul _ _ _ := ext fun _ => add_smul _ _ _ + +instance isCentralScalar [Module S₃ᵐᵒᵖ M₃] [IsCentralScalar S₃ M₃] : + IsCentralScalar S₃ (M →SL[σ₁₃] M₃) where + op_smul_eq_smul _ _ := ext fun _ => op_smul_eq_smul _ _ + +variable (S) [ContinuousAdd N₃] + +/-- The coercion from `M →L[R] M₂` to `M →ₗ[R] M₂`, as a linear map. -/ +@[simps] +def coeLM : (M →L[R] N₃) →ₗ[S] M →ₗ[R] N₃ where + toFun := (↑) + map_add' f g := coe_add f g + map_smul' c f := coe_smul c f + +variable {S} (σ₁₃) + +/-- The coercion from `M →SL[σ] M₂` to `M →ₛₗ[σ] M₂`, as a linear map. -/ +@[simps] +def coeLMₛₗ : (M →SL[σ₁₃] M₃) →ₗ[S₃] M →ₛₗ[σ₁₃] M₃ where + toFun := (↑) + map_add' f g := coe_add f g + map_smul' c f := coe_smul c f + +end SMul + +section SMulRightₗ + +variable {R S T M M₂ : Type*} [Semiring R] [Semiring S] [Semiring T] [Module R S] + [AddCommMonoid M₂] [Module R M₂] [Module S M₂] [IsScalarTower R S M₂] [TopologicalSpace S] + [TopologicalSpace M₂] [ContinuousSMul S M₂] [TopologicalSpace M] [AddCommMonoid M] [Module R M] + [ContinuousAdd M₂] [Module T M₂] [ContinuousConstSMul T M₂] [SMulCommClass R T M₂] + [SMulCommClass S T M₂] + +/-- Given `c : E →L[R] S`, `c.smulRightₗ` is the linear map from `F` to `E →L[R] F` +sending `f` to `fun e => c e • f`. See also `ContinuousLinearMap.smulRightL`. -/ +def smulRightₗ (c : M →L[R] S) : M₂ →ₗ[T] M →L[R] M₂ where + toFun := c.smulRight + map_add' x y := by + ext e + apply smul_add (c e) + map_smul' a x := by + ext e + dsimp + apply smul_comm + +@[simp] +theorem coe_smulRightₗ (c : M →L[R] S) : ⇑(smulRightₗ c : M₂ →ₗ[T] M →L[R] M₂) = c.smulRight := + rfl + +end SMulRightₗ + +section CommRing + +variable {R : Type*} [CommRing R] {M : Type*} [TopologicalSpace M] [AddCommGroup M] {M₂ : Type*} + [TopologicalSpace M₂] [AddCommGroup M₂] {M₃ : Type*} [TopologicalSpace M₃] [AddCommGroup M₃] + [Module R M] [Module R M₂] [Module R M₃] + +variable [TopologicalAddGroup M₂] [ContinuousConstSMul R M₂] + +instance algebra : Algebra R (M₂ →L[R] M₂) := + Algebra.ofModule smul_comp fun _ _ _ => comp_smul _ _ _ + +@[simp] theorem algebraMap_apply (r : R) (m : M₂) : algebraMap R (M₂ →L[R] M₂) r m = r • m := rfl + +end CommRing + +section RestrictScalars + +variable {A M M₂ : Type*} [Ring A] [AddCommGroup M] [AddCommGroup M₂] [Module A M] [Module A M₂] + [TopologicalSpace M] [TopologicalSpace M₂] (R : Type*) [Ring R] [Module R M] [Module R M₂] + [LinearMap.CompatibleSMul M M₂ R A] + +/-- If `A` is an `R`-algebra, then a continuous `A`-linear map can be interpreted as a continuous +`R`-linear map. We assume `LinearMap.CompatibleSMul M M₂ R A` to match assumptions of +`LinearMap.map_smul_of_tower`. -/ +def restrictScalars (f : M →L[A] M₂) : M →L[R] M₂ := + ⟨(f : M →ₗ[A] M₂).restrictScalars R, f.continuous⟩ + +variable {R} + +@[simp] -- @[norm_cast] -- Porting note: This theorem can't be a `norm_cast` theorem. +theorem coe_restrictScalars (f : M →L[A] M₂) : + (f.restrictScalars R : M →ₗ[R] M₂) = (f : M →ₗ[A] M₂).restrictScalars R := + rfl + +@[simp] +theorem coe_restrictScalars' (f : M →L[A] M₂) : ⇑(f.restrictScalars R) = f := + rfl + +@[simp] +theorem restrictScalars_zero : (0 : M →L[A] M₂).restrictScalars R = 0 := + rfl + +section + +variable [TopologicalAddGroup M₂] + +@[simp] +theorem restrictScalars_add (f g : M →L[A] M₂) : + (f + g).restrictScalars R = f.restrictScalars R + g.restrictScalars R := + rfl + +@[simp] +theorem restrictScalars_neg (f : M →L[A] M₂) : (-f).restrictScalars R = -f.restrictScalars R := + rfl + +end + +variable {S : Type*} +variable [Ring S] [Module S M₂] [ContinuousConstSMul S M₂] [SMulCommClass A S M₂] + [SMulCommClass R S M₂] + +@[simp] +theorem restrictScalars_smul (c : S) (f : M →L[A] M₂) : + (c • f).restrictScalars R = c • f.restrictScalars R := + rfl + +variable (A M M₂ R S) +variable [TopologicalAddGroup M₂] + +/-- `ContinuousLinearMap.restrictScalars` as a `LinearMap`. See also +`ContinuousLinearMap.restrictScalarsL`. -/ +def restrictScalarsₗ : (M →L[A] M₂) →ₗ[S] M →L[R] M₂ where + toFun := restrictScalars R + map_add' := restrictScalars_add + map_smul' := restrictScalars_smul + +variable {A M M₂ R S} + +@[simp] +theorem coe_restrictScalarsₗ : ⇑(restrictScalarsₗ A M M₂ R S) = restrictScalars R := + rfl + +end RestrictScalars + +end ContinuousLinearMap + +namespace Submodule + +variable {R : Type*} [Ring R] {M : Type*} [TopologicalSpace M] [AddCommGroup M] [Module R M] + +open ContinuousLinearMap + +/-- A submodule `p` is called *complemented* if there exists a continuous projection `M →ₗ[R] p`. -/ +def ClosedComplemented (p : Submodule R M) : Prop := + ∃ f : M →L[R] p, ∀ x : p, f x = x + +theorem ClosedComplemented.exists_isClosed_isCompl {p : Submodule R M} [T1Space p] + (h : ClosedComplemented p) : + ∃ q : Submodule R M, IsClosed (q : Set M) ∧ IsCompl p q := + Exists.elim h fun f hf => ⟨ker f, isClosed_ker f, LinearMap.isCompl_of_proj hf⟩ + +protected theorem ClosedComplemented.isClosed [TopologicalAddGroup M] [T1Space M] + {p : Submodule R M} (h : ClosedComplemented p) : IsClosed (p : Set M) := by + rcases h with ⟨f, hf⟩ + have : ker (id R M - p.subtypeL.comp f) = p := LinearMap.ker_id_sub_eq_of_proj hf + exact this ▸ isClosed_ker _ + +@[simp] +theorem closedComplemented_bot : ClosedComplemented (⊥ : Submodule R M) := + ⟨0, fun x => by simp only [zero_apply, eq_zero_of_bot_submodule x]⟩ + +@[simp] +theorem closedComplemented_top : ClosedComplemented (⊤ : Submodule R M) := + ⟨(id R M).codRestrict ⊤ fun _x => trivial, fun x => Subtype.ext_iff_val.2 <| by simp⟩ + +end Submodule + +theorem ContinuousLinearMap.closedComplemented_ker_of_rightInverse {R : Type*} [Ring R] + {M : Type*} [TopologicalSpace M] [AddCommGroup M] {M₂ : Type*} [TopologicalSpace M₂] + [AddCommGroup M₂] [Module R M] [Module R M₂] [TopologicalAddGroup M] (f₁ : M →L[R] M₂) + (f₂ : M₂ →L[R] M) (h : Function.RightInverse f₂ f₁) : (ker f₁).ClosedComplemented := + ⟨f₁.projKerOfRightInverse f₂ h, f₁.projKerOfRightInverse_apply_idem f₂ h⟩ diff --git a/Mathlib/Topology/Algebra/Module/LinearMapPiProd.lean b/Mathlib/Topology/Algebra/Module/LinearMapPiProd.lean new file mode 100644 index 0000000000000..7a25e4c44268e --- /dev/null +++ b/Mathlib/Topology/Algebra/Module/LinearMapPiProd.lean @@ -0,0 +1,299 @@ +/- +Copyright (c) 2019 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jan-David Salchow, Sébastien Gouëzel, Jean Lo, Yury Kudryashov, Frédéric Dupuis, + Heather Macbeth +-/ +import Mathlib.Topology.Algebra.Module.LinearMap + +/-! +# Continuous linear maps on products and Pi types +-/ + +assert_not_exists Star.star + +open LinearMap (ker range) +open Topology Filter Pointwise + +universe u v w u' + +namespace ContinuousLinearMap + +section Semiring + +/-! +### Properties that hold for non-necessarily commutative semirings. +-/ + +variable + {R : Type*} [Semiring R] + {M₁ : Type*} [TopologicalSpace M₁] [AddCommMonoid M₁] [Module R M₁] + {M₂ : Type*} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R M₂] + {M₃ : Type*} [TopologicalSpace M₃] [AddCommMonoid M₃] [Module R M₃] + {M₄ : Type*} [TopologicalSpace M₄] [AddCommMonoid M₄] [Module R M₄] + +/-- The cartesian product of two bounded linear maps, as a bounded linear map. -/ +protected def prod (f₁ : M₁ →L[R] M₂) (f₂ : M₁ →L[R] M₃) : + M₁ →L[R] M₂ × M₃ := + ⟨(f₁ : M₁ →ₗ[R] M₂).prod f₂, f₁.2.prod_mk f₂.2⟩ + +@[simp, norm_cast] +theorem coe_prod (f₁ : M₁ →L[R] M₂) (f₂ : M₁ →L[R] M₃) : + (f₁.prod f₂ : M₁ →ₗ[R] M₂ × M₃) = LinearMap.prod f₁ f₂ := + rfl + +@[simp, norm_cast] +theorem prod_apply (f₁ : M₁ →L[R] M₂) (f₂ : M₁ →L[R] M₃) (x : M₁) : + f₁.prod f₂ x = (f₁ x, f₂ x) := + rfl + +section + +variable (R M₁ M₂) + +/-- The left injection into a product is a continuous linear map. -/ +def inl : M₁ →L[R] M₁ × M₂ := + (id R M₁).prod 0 + +/-- The right injection into a product is a continuous linear map. -/ +def inr : M₂ →L[R] M₁ × M₂ := + (0 : M₂ →L[R] M₁).prod (id R M₂) + +end + +@[simp] +theorem inl_apply (x : M₁) : inl R M₁ M₂ x = (x, 0) := + rfl + +@[simp] +theorem inr_apply (x : M₂) : inr R M₁ M₂ x = (0, x) := + rfl + +@[simp, norm_cast] +theorem coe_inl : (inl R M₁ M₂ : M₁ →ₗ[R] M₁ × M₂) = LinearMap.inl R M₁ M₂ := + rfl + +@[simp, norm_cast] +theorem coe_inr : (inr R M₁ M₂ : M₂ →ₗ[R] M₁ × M₂) = LinearMap.inr R M₁ M₂ := + rfl + +@[simp] +theorem ker_prod (f : M₁ →L[R] M₂) (g : M₁ →L[R] M₃) : + ker (f.prod g) = ker f ⊓ ker g := + LinearMap.ker_prod (f : M₁ →ₗ[R] M₂) (g : M₁ →ₗ[R] M₃) + +variable (R M₁ M₂) + +/-- `Prod.fst` as a `ContinuousLinearMap`. -/ +def fst : M₁ × M₂ →L[R] M₁ where + cont := continuous_fst + toLinearMap := LinearMap.fst R M₁ M₂ + +/-- `Prod.snd` as a `ContinuousLinearMap`. -/ +def snd : M₁ × M₂ →L[R] M₂ where + cont := continuous_snd + toLinearMap := LinearMap.snd R M₁ M₂ + +variable {R M₁ M₂} + +@[simp, norm_cast] +theorem coe_fst : ↑(fst R M₁ M₂) = LinearMap.fst R M₁ M₂ := + rfl + +@[simp, norm_cast] +theorem coe_fst' : ⇑(fst R M₁ M₂) = Prod.fst := + rfl + +@[simp, norm_cast] +theorem coe_snd : ↑(snd R M₁ M₂) = LinearMap.snd R M₁ M₂ := + rfl + +@[simp, norm_cast] +theorem coe_snd' : ⇑(snd R M₁ M₂) = Prod.snd := + rfl + +@[simp] +theorem fst_prod_snd : (fst R M₁ M₂).prod (snd R M₁ M₂) = id R (M₁ × M₂) := + ext fun ⟨_x, _y⟩ => rfl + +@[simp] +theorem fst_comp_prod (f : M₁ →L[R] M₂) (g : M₁ →L[R] M₃) : + (fst R M₂ M₃).comp (f.prod g) = f := + ext fun _x => rfl + +@[simp] +theorem snd_comp_prod (f : M₁ →L[R] M₂) (g : M₁ →L[R] M₃) : + (snd R M₂ M₃).comp (f.prod g) = g := + ext fun _x => rfl + +/-- `Prod.map` of two continuous linear maps. -/ +def prodMap (f₁ : M₁ →L[R] M₂) (f₂ : M₃ →L[R] M₄) : + M₁ × M₃ →L[R] M₂ × M₄ := + (f₁.comp (fst R M₁ M₃)).prod (f₂.comp (snd R M₁ M₃)) + +@[simp, norm_cast] +theorem coe_prodMap (f₁ : M₁ →L[R] M₂) + (f₂ : M₃ →L[R] M₄) : ↑(f₁.prodMap f₂) = (f₁ : M₁ →ₗ[R] M₂).prodMap (f₂ : M₃ →ₗ[R] M₄) := + rfl + +@[simp, norm_cast] +theorem coe_prodMap' (f₁ : M₁ →L[R] M₂) + (f₂ : M₃ →L[R] M₄) : ⇑(f₁.prodMap f₂) = Prod.map f₁ f₂ := + rfl + +/-- The continuous linear map given by `(x, y) ↦ f₁ x + f₂ y`. -/ +def coprod [ContinuousAdd M₃] (f₁ : M₁ →L[R] M₃) + (f₂ : M₂ →L[R] M₃) : M₁ × M₂ →L[R] M₃ := + ⟨LinearMap.coprod f₁ f₂, (f₁.cont.comp continuous_fst).add (f₂.cont.comp continuous_snd)⟩ + +@[norm_cast, simp] +theorem coe_coprod [ContinuousAdd M₃] (f₁ : M₁ →L[R] M₃) + (f₂ : M₂ →L[R] M₃) : (f₁.coprod f₂ : M₁ × M₂ →ₗ[R] M₃) = LinearMap.coprod f₁ f₂ := + rfl + +@[simp] +theorem coprod_apply [ContinuousAdd M₃] (f₁ : M₁ →L[R] M₃) + (f₂ : M₂ →L[R] M₃) (x) : f₁.coprod f₂ x = f₁ x.1 + f₂ x.2 := + rfl + +theorem range_coprod [ContinuousAdd M₃] (f₁ : M₁ →L[R] M₃) + (f₂ : M₂ →L[R] M₃) : range (f₁.coprod f₂) = range f₁ ⊔ range f₂ := + LinearMap.range_coprod _ _ + +theorem comp_fst_add_comp_snd [ContinuousAdd M₃] (f : M₁ →L[R] M₃) + (g : M₂ →L[R] M₃) : + f.comp (ContinuousLinearMap.fst R M₁ M₂) + g.comp (ContinuousLinearMap.snd R M₁ M₂) = + f.coprod g := + rfl + +theorem coprod_inl_inr [ContinuousAdd M₁] [ContinuousAdd M₂] : + (ContinuousLinearMap.inl R M₁ M₂).coprod (ContinuousLinearMap.inr R M₁ M₂) = + ContinuousLinearMap.id R (M₁ × M₂) := by + apply coe_injective; apply LinearMap.coprod_inl_inr + +end Semiring + +section Pi + +variable {R : Type*} [Semiring R] {M : Type*} [TopologicalSpace M] [AddCommMonoid M] [Module R M] + {M₂ : Type*} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R M₂] {ι : Type*} {φ : ι → Type*} + [∀ i, TopologicalSpace (φ i)] [∀ i, AddCommMonoid (φ i)] [∀ i, Module R (φ i)] + +/-- `pi` construction for continuous linear functions. From a family of continuous linear functions +it produces a continuous linear function into a family of topological modules. -/ +def pi (f : ∀ i, M →L[R] φ i) : M →L[R] ∀ i, φ i := + ⟨LinearMap.pi fun i => f i, continuous_pi fun i => (f i).continuous⟩ + +@[simp] +theorem coe_pi' (f : ∀ i, M →L[R] φ i) : ⇑(pi f) = fun c i => f i c := + rfl + +@[simp] +theorem coe_pi (f : ∀ i, M →L[R] φ i) : (pi f : M →ₗ[R] ∀ i, φ i) = LinearMap.pi fun i => f i := + rfl + +theorem pi_apply (f : ∀ i, M →L[R] φ i) (c : M) (i : ι) : pi f c i = f i c := + rfl + +theorem pi_eq_zero (f : ∀ i, M →L[R] φ i) : pi f = 0 ↔ ∀ i, f i = 0 := by + simp only [ContinuousLinearMap.ext_iff, pi_apply, funext_iff] + exact forall_swap + +theorem pi_zero : pi (fun _ => 0 : ∀ i, M →L[R] φ i) = 0 := + ext fun _ => rfl + +theorem pi_comp (f : ∀ i, M →L[R] φ i) (g : M₂ →L[R] M) : + (pi f).comp g = pi fun i => (f i).comp g := + rfl + +/-- The projections from a family of topological modules are continuous linear maps. -/ +def proj (i : ι) : (∀ i, φ i) →L[R] φ i := + ⟨LinearMap.proj i, continuous_apply _⟩ + +@[simp] +theorem proj_apply (i : ι) (b : ∀ i, φ i) : (proj i : (∀ i, φ i) →L[R] φ i) b = b i := + rfl + +theorem proj_pi (f : ∀ i, M₂ →L[R] φ i) (i : ι) : (proj i).comp (pi f) = f i := + ext fun _c => rfl + +theorem iInf_ker_proj : (⨅ i, ker (proj i : (∀ i, φ i) →L[R] φ i) : Submodule R (∀ i, φ i)) = ⊥ := + LinearMap.iInf_ker_proj + +variable (R φ) + +/-- Given a function `f : α → ι`, it induces a continuous linear function by right composition on +product types. For `f = Subtype.val`, this corresponds to forgetting some set of variables. -/ +def _root_.Pi.compRightL {α : Type*} (f : α → ι) : ((i : ι) → φ i) →L[R] ((i : α) → φ (f i)) where + toFun := fun v i ↦ v (f i) + map_add' := by intros; ext; simp + map_smul' := by intros; ext; simp + cont := by continuity + +@[simp] lemma _root_.Pi.compRightL_apply {α : Type*} (f : α → ι) (v : (i : ι) → φ i) (i : α) : + Pi.compRightL R φ f v i = v (f i) := rfl + +end Pi + +section Ring + +variable {R : Type*} [Ring R] + {M : Type*} [TopologicalSpace M] [AddCommGroup M] [Module R M] + {M₂ : Type*} [TopologicalSpace M₂] [AddCommGroup M₂] [Module R M₂] + {M₃ : Type*} [TopologicalSpace M₃] [AddCommGroup M₃] [Module R M₃] + +theorem range_prod_eq {f : M →L[R] M₂} {g : M →L[R] M₃} (h : ker f ⊔ ker g = ⊤) : + range (f.prod g) = (range f).prod (range g) := + LinearMap.range_prod_eq h + +theorem ker_prod_ker_le_ker_coprod [ContinuousAdd M₃] (f : M →L[R] M₃) (g : M₂ →L[R] M₃) : + (LinearMap.ker f).prod (LinearMap.ker g) ≤ LinearMap.ker (f.coprod g) := + LinearMap.ker_prod_ker_le_ker_coprod f.toLinearMap g.toLinearMap + +theorem ker_coprod_of_disjoint_range [ContinuousAdd M₃] (f : M →L[R] M₃) (g : M₂ →L[R] M₃) + (hd : Disjoint (range f) (range g)) : + LinearMap.ker (f.coprod g) = (LinearMap.ker f).prod (LinearMap.ker g) := + LinearMap.ker_coprod_of_disjoint_range f.toLinearMap g.toLinearMap hd + +end Ring + +section SMul + +variable + {R : Type*} [Semiring R] + {M : Type*} [TopologicalSpace M] [AddCommMonoid M] [Module R M] + {M₂ : Type*} [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R M₂] + {M₃ : Type*} [TopologicalSpace M₃] [AddCommMonoid M₃] [Module R M₃] + +/-- `ContinuousLinearMap.prod` as an `Equiv`. -/ +@[simps apply] +def prodEquiv : (M →L[R] M₂) × (M →L[R] M₃) ≃ (M →L[R] M₂ × M₃) where + toFun f := f.1.prod f.2 + invFun f := ⟨(fst _ _ _).comp f, (snd _ _ _).comp f⟩ + left_inv f := by ext <;> rfl + right_inv f := by ext <;> rfl + +theorem prod_ext_iff {f g : M × M₂ →L[R] M₃} : + f = g ↔ f.comp (inl _ _ _) = g.comp (inl _ _ _) ∧ f.comp (inr _ _ _) = g.comp (inr _ _ _) := by + simp only [← coe_inj, LinearMap.prod_ext_iff] + rfl + +@[ext] +theorem prod_ext {f g : M × M₂ →L[R] M₃} (hl : f.comp (inl _ _ _) = g.comp (inl _ _ _)) + (hr : f.comp (inr _ _ _) = g.comp (inr _ _ _)) : f = g := + prod_ext_iff.2 ⟨hl, hr⟩ + +variable (S : Type*) [Semiring S] + [Module S M₂] [ContinuousAdd M₂] [SMulCommClass R S M₂] [ContinuousConstSMul S M₂] + [Module S M₃] [ContinuousAdd M₃] [SMulCommClass R S M₃] [ContinuousConstSMul S M₃] + +/-- `ContinuousLinearMap.prod` as a `LinearEquiv`. -/ +@[simps apply] +def prodₗ : ((M →L[R] M₂) × (M →L[R] M₃)) ≃ₗ[S] M →L[R] M₂ × M₃ := + { prodEquiv with + map_add' := fun _f _g => rfl + map_smul' := fun _c _f => rfl } + +end SMul + +end ContinuousLinearMap diff --git a/Mathlib/Topology/Algebra/Module/ModuleTopology.lean b/Mathlib/Topology/Algebra/Module/ModuleTopology.lean index e5366ba9ea100..e19beebd3a3cd 100644 --- a/Mathlib/Topology/Algebra/Module/ModuleTopology.lean +++ b/Mathlib/Topology/Algebra/Module/ModuleTopology.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Kevin Buzzard. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kevin Buzzard, Will Sawin -/ -import Mathlib.Topology.Algebra.Module.Basic +import Mathlib.Topology.Algebra.Module.Equiv /-! # A "module topology" for modules over a topological ring diff --git a/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean b/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean index f7e4876c83173..65348675488a9 100644 --- a/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean @@ -3,8 +3,8 @@ Copyright (c) 2020 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import Mathlib.Topology.Algebra.Module.Basic import Mathlib.LinearAlgebra.Multilinear.Basic +import Mathlib.Topology.Algebra.Module.LinearMapPiProd /-! # Continuous multilinear maps diff --git a/Mathlib/Topology/Algebra/Module/Star.lean b/Mathlib/Topology/Algebra/Module/Star.lean index 43c8205a370e1..11d8032d76fb5 100644 --- a/Mathlib/Topology/Algebra/Module/Star.lean +++ b/Mathlib/Topology/Algebra/Module/Star.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser, Frédéric Dupuis -/ import Mathlib.Algebra.Star.Module -import Mathlib.Topology.Algebra.Module.Basic +import Mathlib.Topology.Algebra.Module.Equiv import Mathlib.Topology.Algebra.Star /-! diff --git a/Mathlib/Topology/Algebra/Module/StrongTopology.lean b/Mathlib/Topology/Algebra/Module/StrongTopology.lean index 9da5fa734ddb3..543fab13c9fbb 100644 --- a/Mathlib/Topology/Algebra/Module/StrongTopology.lean +++ b/Mathlib/Topology/Algebra/Module/StrongTopology.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Anatole Dedecker. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anatole Dedecker, Yury Kudryashov -/ +import Mathlib.Topology.Algebra.Module.Equiv import Mathlib.Topology.Algebra.Module.UniformConvergence import Mathlib.Topology.Algebra.SeparationQuotient.Section import Mathlib.Topology.Hom.ContinuousEvalConst diff --git a/Mathlib/Topology/Algebra/Module/WeakDual.lean b/Mathlib/Topology/Algebra/Module/WeakDual.lean index d2bdbb2408760..cc9ff04926313 100644 --- a/Mathlib/Topology/Algebra/Module/WeakDual.lean +++ b/Mathlib/Topology/Algebra/Module/WeakDual.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kalle Kytölä, Moritz Doll -/ import Mathlib.LinearAlgebra.BilinearMap -import Mathlib.Topology.Algebra.Module.Basic +import Mathlib.Topology.Algebra.Module.LinearMap import Mathlib.Topology.Algebra.Module.WeakBilin /-! diff --git a/Mathlib/Topology/Algebra/Nonarchimedean/Bases.lean b/Mathlib/Topology/Algebra/Nonarchimedean/Bases.lean index 6fd78154fd5a4..19fcae3506dc6 100644 --- a/Mathlib/Topology/Algebra/Nonarchimedean/Bases.lean +++ b/Mathlib/Topology/Algebra/Nonarchimedean/Bases.lean @@ -3,9 +3,10 @@ Copyright (c) 2021 Patrick Massot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot -/ -import Mathlib.Topology.Algebra.Nonarchimedean.Basic -import Mathlib.Topology.Algebra.FilterBasis +import Mathlib.Algebra.Algebra.Basic import Mathlib.Algebra.Module.Submodule.Pointwise +import Mathlib.Topology.Algebra.FilterBasis +import Mathlib.Topology.Algebra.Nonarchimedean.Basic /-! # Neighborhood bases for non-archimedean rings and modules diff --git a/Mathlib/Topology/Algebra/SeparationQuotient/Basic.lean b/Mathlib/Topology/Algebra/SeparationQuotient/Basic.lean index 6790b8c087127..14e8443fb52c5 100644 --- a/Mathlib/Topology/Algebra/SeparationQuotient/Basic.lean +++ b/Mathlib/Topology/Algebra/SeparationQuotient/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import Mathlib.Topology.Algebra.Module.Basic +import Mathlib.Topology.Algebra.Module.LinearMap /-! # Algebraic operations on `SeparationQuotient` diff --git a/Mathlib/Topology/Algebra/UniformFilterBasis.lean b/Mathlib/Topology/Algebra/UniformFilterBasis.lean index 055353936b31b..7d14667ebd029 100644 --- a/Mathlib/Topology/Algebra/UniformFilterBasis.lean +++ b/Mathlib/Topology/Algebra/UniformFilterBasis.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot -/ import Mathlib.Topology.Algebra.FilterBasis +import Mathlib.Topology.Algebra.UniformGroup.Defs /-! # Uniform properties of neighborhood bases in topological algebra diff --git a/Mathlib/Topology/ContinuousMap/Algebra.lean b/Mathlib/Topology/ContinuousMap/Algebra.lean index 58c98d954d09d..7b2984309d3c0 100644 --- a/Mathlib/Topology/ContinuousMap/Algebra.lean +++ b/Mathlib/Topology/ContinuousMap/Algebra.lean @@ -6,8 +6,8 @@ Authors: Kim Morrison, Nicolò Cavalleri import Mathlib.Algebra.Algebra.Pi import Mathlib.Algebra.Algebra.Subalgebra.Basic import Mathlib.Tactic.FieldSimp -import Mathlib.Topology.Algebra.Module.Basic import Mathlib.Topology.Algebra.InfiniteSum.Basic +import Mathlib.Topology.Algebra.Module.LinearMap import Mathlib.Topology.Algebra.Ring.Basic import Mathlib.Topology.UniformSpace.CompactConvergence diff --git a/Mathlib/Topology/Instances/RealVectorSpace.lean b/Mathlib/Topology/Instances/RealVectorSpace.lean index 5d4327d792021..0fd3382672de6 100644 --- a/Mathlib/Topology/Instances/RealVectorSpace.lean +++ b/Mathlib/Topology/Instances/RealVectorSpace.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 -/ -import Mathlib.Topology.Algebra.Module.Basic +import Mathlib.Topology.Algebra.Module.Equiv import Mathlib.Topology.Instances.Rat import Mathlib.Algebra.Module.Rat diff --git a/Mathlib/Topology/Instances/TrivSqZeroExt.lean b/Mathlib/Topology/Instances/TrivSqZeroExt.lean index 5a6add1956758..3c28a5b432624 100644 --- a/Mathlib/Topology/Instances/TrivSqZeroExt.lean +++ b/Mathlib/Topology/Instances/TrivSqZeroExt.lean @@ -5,7 +5,7 @@ Authors: Eric Wieser -/ import Mathlib.Algebra.TrivSqZeroExt import Mathlib.Topology.Algebra.InfiniteSum.Basic -import Mathlib.Topology.Algebra.Module.Basic +import Mathlib.Topology.Algebra.Module.LinearMapPiProd /-! # Topology on `TrivSqZeroExt R M` From c91dd0e151d3f0b6755d35119cd7943a516addb8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Thu, 19 Dec 2024 21:50:31 +0000 Subject: [PATCH 793/829] refactor(CategoryTheory/Sites): define Functor.sheafPullback as a left adjoint (#19516) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The pullback of sheaves was previously constructed as a composition of functors involving a left Kan extension. We redefine it as an abstract left adjoint of the pushforward functor when it exists, and in the most favourable case, we provide an isomorphism with the previous construction. (This change shall be needed when we want to take pullback of étale sheaves.) Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- Mathlib/CategoryTheory/Sites/Pullback.lean | 100 ++++++++++++++++----- Mathlib/Topology/Sheaves/Functors.lean | 6 +- Mathlib/Topology/Sheaves/Stalks.lean | 1 - 3 files changed, 79 insertions(+), 28 deletions(-) diff --git a/Mathlib/CategoryTheory/Sites/Pullback.lean b/Mathlib/CategoryTheory/Sites/Pullback.lean index ccf9e1aa3527c..26b39bd096bdb 100644 --- a/Mathlib/CategoryTheory/Sites/Pullback.lean +++ b/Mathlib/CategoryTheory/Sites/Pullback.lean @@ -13,8 +13,10 @@ import Mathlib.CategoryTheory.Sites.LeftExact ## Main definitions -* `CategoryTheory.Functor.sheafPullback`: the functor `Sheaf J A ⥤ Sheaf K A` obtained -as an extension of a functor `G : C ⥤ D` between the underlying categories. +* `CategoryTheory.Functor.sheafPullback`: when `G : C ⥤ D` is a continuous functor + between sites (for topologies `J` on `C` and `K` on `D`) such that the functor + `G.sheafPushforwardContinuous A J K : Sheaf K A ⥤ Sheaf J A` has a left adjoint, + this is the pullback functor defined as a chosen left adjoint. * `CategoryTheory.Functor.sheafAdjunctionContinuous`: the adjunction `G.sheafPullback A J K ⊣ G.sheafPushforwardContinuous A J K` when the functor @@ -25,42 +27,94 @@ sense of SGA 4 IV 4.9. -/ -universe v₁ u₁ +universe v₁ v₂ v₃ u₁ u₂ u₃ noncomputable section -open CategoryTheory.Limits +namespace CategoryTheory.Functor -namespace CategoryTheory +open Limits -variable {C : Type v₁} [SmallCategory C] {D : Type v₁} [SmallCategory D] (G : C ⥤ D) -variable (A : Type u₁) [Category.{v₁} A] -variable (J : GrothendieckTopology C) (K : GrothendieckTopology D) +section GeneralUniverses --- The assumptions so that we have sheafification -variable [ConcreteCategory.{v₁} A] [PreservesLimits (forget A)] [HasColimits A] [HasLimits A] -variable [PreservesFilteredColimits (forget A)] [(forget A).ReflectsIsomorphisms] +variable {C : Type u₂} [Category.{v₂} C] {D : Type u₃} [Category.{v₃} D] (G : C ⥤ D) + (A : Type u₁) [Category.{v₁} A] + (J : GrothendieckTopology C) (K : GrothendieckTopology D) + [Functor.IsContinuous.{v₁} G J K] -attribute [local instance] reflectsLimits_of_reflectsIsomorphisms +section -instance {X : C} : IsCofiltered (J.Cover X) := - inferInstance +variable [(G.sheafPushforwardContinuous A J K).IsRightAdjoint] /-- The pullback functor `Sheaf J A ⥤ Sheaf K A` associated to a functor `G : C ⥤ D` in the same direction as `G`. -/ -@[simps!] -def Functor.sheafPullback : Sheaf J A ⥤ Sheaf K A := +def sheafPullback : Sheaf J A ⥤ Sheaf K A := + (G.sheafPushforwardContinuous A J K).leftAdjoint + +/-- The pullback functor is left adjoint to the pushforward functor. -/ +def sheafAdjunctionContinuous : + G.sheafPullback A J K ⊣ G.sheafPushforwardContinuous A J K := + Adjunction.ofIsRightAdjoint (G.sheafPushforwardContinuous A J K) + +end + +namespace sheafPullbackConstruction + +variable [∀ (F : Cᵒᵖ ⥤ A), G.op.HasLeftKanExtension F] + +/-- Construction of the pullback of sheaves using a left Kan extension. -/ +def sheafPullback [HasWeakSheafify K A] : Sheaf J A ⥤ Sheaf K A := sheafToPresheaf J A ⋙ G.op.lan ⋙ presheafToSheaf K A -instance [RepresentablyFlat G] : PreservesFiniteLimits (G.sheafPullback A J K) := by +/-- The constructed `sheafPullback G A J K` is left adjoint +to `G.sheafPushforwardContinuous A J K`. -/ +def sheafAdjunctionContinuous [Functor.IsContinuous.{v₁} G J K] [HasWeakSheafify K A] : + sheafPullback G A J K ⊣ G.sheafPushforwardContinuous A J K := + ((G.op.lanAdjunction A).comp (sheafificationAdjunction K A)).restrictFullyFaithful + (fullyFaithfulSheafToPresheaf J A) (Functor.FullyFaithful.id _) (Iso.refl _) (Iso.refl _) + +instance [HasWeakSheafify K A] : + (G.sheafPushforwardContinuous A J K).IsRightAdjoint := + (sheafAdjunctionContinuous G A J K).isRightAdjoint + +/-- The constructed pullback of sheaves is isomorphic to the abstract one. -/ +def sheafPullbackIso [HasWeakSheafify K A] : + Functor.sheafPullback G A J K ≅ sheafPullback G A J K := + Adjunction.leftAdjointUniq (Functor.sheafAdjunctionContinuous G A J K) + (sheafAdjunctionContinuous G A J K) + +variable [RepresentablyFlat G] [HasSheafify K A] [HasSheafify J A] + [PreservesFiniteLimits (G.op.lan : (_ ⥤ _ ⥤ A))] + +instance : PreservesFiniteLimits (sheafPullback G A J K) := by have : PreservesFiniteLimits (G.op.lan ⋙ presheafToSheaf K A) := comp_preservesFiniteLimits _ _ apply comp_preservesFiniteLimits -/-- The pullback functor is left adjoint to the pushforward functor. -/ -def Functor.sheafAdjunctionContinuous [Functor.IsContinuous.{v₁} G J K] : - G.sheafPullback A J K ⊣ G.sheafPushforwardContinuous A J K := - ((G.op.lanAdjunction A).comp (sheafificationAdjunction K A)).restrictFullyFaithful - (fullyFaithfulSheafToPresheaf J A) (Functor.FullyFaithful.id _) (Iso.refl _) (Iso.refl _) +instance preservesFiniteLimits : PreservesFiniteLimits (Functor.sheafPullback G A J K) := + preservesFiniteLimits_of_natIso (sheafPullbackIso G A J K).symm + +end sheafPullbackConstruction + +end GeneralUniverses + +namespace SmallCategories + +variable {C : Type v₁} [SmallCategory C] {D : Type v₁} [SmallCategory D] (G : C ⥤ D) + (A : Type u₁) [Category.{v₁} A] + (J : GrothendieckTopology C) (K : GrothendieckTopology D) + +-- The favourable assumptions under which we have sheafification +variable [ConcreteCategory.{v₁} A] [PreservesLimits (forget A)] [HasColimits A] [HasLimits A] + [PreservesFilteredColimits (forget A)] [(forget A).ReflectsIsomorphisms] + [Functor.IsContinuous.{v₁} G J K] + +example : (G.sheafPushforwardContinuous A J K).IsRightAdjoint := inferInstance + +attribute [local instance] reflectsLimits_of_reflectsIsomorphisms in +instance [RepresentablyFlat G] : PreservesFiniteLimits (G.sheafPullback A J K) := by + apply sheafPullbackConstruction.preservesFiniteLimits + +end SmallCategories -end CategoryTheory +end CategoryTheory.Functor diff --git a/Mathlib/Topology/Sheaves/Functors.lean b/Mathlib/Topology/Sheaves/Functors.lean index d246fbdbe9fcc..4bf77afd2c82a 100644 --- a/Mathlib/Topology/Sheaves/Functors.lean +++ b/Mathlib/Topology/Sheaves/Functors.lean @@ -86,15 +86,13 @@ The pullback functor. def pullback (f : X ⟶ Y) : Y.Sheaf A ⥤ X.Sheaf A := (Opens.map f).sheafPullback _ _ _ -lemma pullback_eq (f : X ⟶ Y) : - pullback A f = forget A Y ⋙ Presheaf.pullback A f ⋙ presheafToSheaf _ _ := rfl - /-- The pullback of a sheaf is isomorphic (actually definitionally equal) to the sheafification of the pullback as a presheaf. -/ def pullbackIso (f : X ⟶ Y) : - pullback A f ≅ forget A Y ⋙ Presheaf.pullback A f ⋙ presheafToSheaf _ _ := Iso.refl _ + pullback A f ≅ forget A Y ⋙ Presheaf.pullback A f ⋙ presheafToSheaf _ _ := + Functor.sheafPullbackConstruction.sheafPullbackIso _ _ _ _ /-- The adjunction between pullback and pushforward for sheaves on topological spaces. -/ def pullbackPushforwardAdjunction (f : X ⟶ Y) : diff --git a/Mathlib/Topology/Sheaves/Stalks.lean b/Mathlib/Topology/Sheaves/Stalks.lean index 8c4b279f06b5d..28c58d23c6a6d 100644 --- a/Mathlib/Topology/Sheaves/Stalks.lean +++ b/Mathlib/Topology/Sheaves/Stalks.lean @@ -3,7 +3,6 @@ Copyright (c) 2019 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison, Justus Springer -/ -import Mathlib.CategoryTheory.Sites.Pullback import Mathlib.Topology.Category.TopCat.OpenNhds import Mathlib.Topology.Sheaves.SheafCondition.UniqueGluing From 0065e8d93b880f81391d0c21f3aa49a06f80ef87 Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Fri, 20 Dec 2024 00:42:07 +0000 Subject: [PATCH 794/829] feat(AlgebraicTopology/HomotopyCat): SSet.hoFunctor (#16783) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This defines `SSet.hoFunctor`, which constructs a category from a simplicial set. Co-Authored-By: Emily Riehl , Pietro Monticone <38562595+pitmonticone@users.noreply.github.com>, and Joël Riou Co-authored-by: Emily Riehl Co-authored-by: emilyriehl Co-authored-by: Mario Carneiro --- Mathlib.lean | 1 + .../SimplicialObject/Basic.lean | 2 +- .../SimplicialSet/Basic.lean | 14 + .../SimplicialSet/HomotopyCat.lean | 341 ++++++++++++++++++ Mathlib/CategoryTheory/Category/Quiv.lean | 65 ++++ Mathlib/CategoryTheory/Category/ReflQuiv.lean | 26 +- Mathlib/CategoryTheory/Quotient.lean | 34 ++ Mathlib/Combinatorics/Quiver/Basic.lean | 28 ++ Mathlib/Combinatorics/Quiver/Prefunctor.lean | 29 ++ Mathlib/Combinatorics/Quiver/ReflQuiver.lean | 16 + 10 files changed, 553 insertions(+), 3 deletions(-) create mode 100644 Mathlib/AlgebraicTopology/SimplicialSet/HomotopyCat.lean diff --git a/Mathlib.lean b/Mathlib.lean index f04515a65fed1..88d19be469fa5 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1061,6 +1061,7 @@ import Mathlib.AlgebraicTopology.SimplicialObject.Basic import Mathlib.AlgebraicTopology.SimplicialObject.Coskeletal import Mathlib.AlgebraicTopology.SimplicialSet.Basic import Mathlib.AlgebraicTopology.SimplicialSet.Coskeletal +import Mathlib.AlgebraicTopology.SimplicialSet.HomotopyCat import Mathlib.AlgebraicTopology.SimplicialSet.KanComplex import Mathlib.AlgebraicTopology.SimplicialSet.Monoidal import Mathlib.AlgebraicTopology.SimplicialSet.Nerve diff --git a/Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean b/Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean index 777aac3b212a6..d8a661f4056b0 100644 --- a/Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean +++ b/Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean @@ -50,7 +50,7 @@ set_option quotPrecheck false in /-- `X _[n]` denotes the `n`th-term of the simplicial object X -/ scoped[Simplicial] notation3:1000 X " _[" n "]" => - (X : CategoryTheory.SimplicialObject _).obj (Opposite.op (SimplexCategory.mk n)) + (X : CategoryTheory.SimplicialObject _).obj (Opposite.op (SimplexCategory.mk n)) open Simplicial diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean index bc6345190d454..bddd445599fef 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean @@ -519,6 +519,20 @@ lemma δ_comp_σ_of_gt'_apply {n} {i : Fin (n + 3)} {j : Fin (n + 2)} (H : j.suc lemma σ_comp_σ_apply {n} {i j : Fin (n + 1)} (H : i ≤ j) (x : S _[n]) : S.σ i.castSucc (S.σ j x) = S.σ j.succ (S.σ i x) := congr_fun (S.σ_comp_σ H) x +variable {T : SSet} (f : S ⟶ T) + +open Opposite + +lemma δ_naturality_apply {n : ℕ} (i : Fin (n + 2)) (x : S _[n + 1]) : + f.app (op [n]) (S.δ i x) = T.δ i (f.app (op [n + 1]) x) := by + show (S.δ i ≫ f.app (op [n])) x = (f.app (op [n + 1]) ≫ T.δ i) x + exact congr_fun (SimplicialObject.δ_naturality f i) x + +lemma σ_naturality_apply {n : ℕ} (i : Fin (n + 1)) (x : S _[n]) : + f.app (op [n + 1]) (S.σ i x) = T.σ i (f.app (op [n]) x) := by + show (S.σ i ≫ f.app (op [n + 1])) x = (f.app (op [n]) ≫ T.σ i) x + exact congr_fun (SimplicialObject.σ_naturality f i) x + end applications end SSet diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/HomotopyCat.lean b/Mathlib/AlgebraicTopology/SimplicialSet/HomotopyCat.lean new file mode 100644 index 0000000000000..46f1b03762397 --- /dev/null +++ b/Mathlib/AlgebraicTopology/SimplicialSet/HomotopyCat.lean @@ -0,0 +1,341 @@ +/- +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, Joël Riou +-/ + +import Mathlib.AlgebraicTopology.SimplicialSet.Coskeletal +import Mathlib.CategoryTheory.Category.ReflQuiv +import Mathlib.Combinatorics.Quiver.ReflQuiver + + +/-! + +# The homotopy category of a simplicial set + +The homotopy category of a simplicial set is defined as a quotient of the free category on its +underlying reflexive quiver (equivalently its one truncation). The quotient imposes an additional +hom relation on this free category, asserting that `f ≫ g = h` whenever `f`, `g`, and `h` are +respectively the 2nd, 0th, and 1st faces of a 2-simplex. + +In fact, the associated functor + +`SSet.hoFunctor : SSet.{u} ⥤ Cat.{u, u} := SSet.truncation 2 ⋙ SSet.hoFunctor₂` + +is defined by first restricting from simplicial sets to 2-truncated simplicial sets (throwing away +the data that is not used for the construction of the homotopy category) and then composing with an +analogously defined `SSet.hoFunctor₂ : SSet.Truncated.{u} 2 ⥤ Cat.{u,u}` implemented relative to +the syntax of the 2-truncated simplex category. + +TODO: Future work will show the functor `SSet.hoFunctor` to be left adjoint to the nerve by +providing an analogous decomposition of the nerve functor, made by possible by the fact that nerves +of categories are 2-coskeletal, and then composing a pair of adjunctions, which factor through the +category of 2-truncated simplicial sets. +-/ + +namespace SSet +open CategoryTheory Category Limits Functor Opposite Simplicial Nerve +universe v u + +section + +local macro:1000 (priority := high) X:term " _[" n:term "]₂" : term => + `(($X : SSet.Truncated 2).obj (Opposite.op ⟨SimplexCategory.mk $n, by decide⟩)) + +set_option quotPrecheck false +local macro:max (priority := high) "[" n:term "]₂" : term => + `((⟨SimplexCategory.mk $n, by decide⟩ : SimplexCategory.Truncated 2)) + +/-- A 2-truncated simplicial set `S` has an underlying refl quiver with `S _[0]₂` as its underlying +type. -/ +def OneTruncation₂ (S : SSet.Truncated 2) := S _[0]₂ + +/-- Abbreviations for face maps in the 2-truncated simplex category. -/ +abbrev δ₂ {n} (i : Fin (n + 2)) (hn := by decide) (hn' := by decide) : + (⟨[n], hn⟩ : SimplexCategory.Truncated 2) ⟶ ⟨[n + 1], hn'⟩ := SimplexCategory.δ i + +/-- Abbreviations for degeneracy maps in the 2-truncated simplex category. -/ +abbrev σ₂ {n} (i : Fin (n + 1)) (hn := by decide) (hn' := by decide) : + (⟨[n+1], hn⟩ : SimplexCategory.Truncated 2) ⟶ ⟨[n], hn'⟩ := SimplexCategory.σ i + +@[reassoc (attr := simp)] +lemma δ₂_zero_comp_σ₂_zero : δ₂ (0 : Fin 2) ≫ σ₂ 0 = 𝟙 _ := SimplexCategory.δ_comp_σ_self + +@[reassoc (attr := simp)] +lemma δ₂_one_comp_σ₂_zero : δ₂ (1 : Fin 2) ≫ σ₂ 0 = 𝟙 _ := SimplexCategory.δ_comp_σ_succ + +/-- The hom-types of the refl quiver underlying a simplicial set `S` are types of edges in `S _[1]₂` +together with source and target equalities. -/ +@[ext] +structure OneTruncation₂.Hom {S : SSet.Truncated 2} (X Y : OneTruncation₂ S) where + /-- An arrow in `OneTruncation₂.Hom X Y` includes the data of a 1-simplex. -/ + edge : S _[1]₂ + /-- An arrow in `OneTruncation₂.Hom X Y` includes a source equality. -/ + src_eq : S.map (δ₂ 1).op edge = X + /-- An arrow in `OneTruncation₂.Hom X Y` includes a target equality. -/ + tgt_eq : S.map (δ₂ 0).op edge = Y + +/-- A 2-truncated simplicial set `S` has an underlying refl quiver `SSet.OneTruncation₂ S`. -/ +instance (S : SSet.Truncated 2) : ReflQuiver (OneTruncation₂ S) where + Hom X Y := SSet.OneTruncation₂.Hom X Y + id X := + { edge := S.map (SSet.σ₂ (n := 0) 0).op X + src_eq := by + simp only [← FunctorToTypes.map_comp_apply, ← op_comp, δ₂_one_comp_σ₂_zero, + op_id, FunctorToTypes.map_id_apply] + tgt_eq := by + simp only [← FunctorToTypes.map_comp_apply, ← op_comp, δ₂_zero_comp_σ₂_zero, + op_id, FunctorToTypes.map_id_apply] } + +@[simp] +lemma OneTruncation₂.id_edge {S : SSet.Truncated 2} (X : OneTruncation₂ S) : + OneTruncation₂.Hom.edge (𝟙rq X) = S.map (SSet.σ₂ 0).op X := rfl + +/-- The functor that carries a 2-truncated simplicial set to its underlying refl quiver. -/ +@[simps] +def oneTruncation₂ : SSet.Truncated.{u} 2 ⥤ ReflQuiv.{u, u} where + obj S := ReflQuiv.of (OneTruncation₂ S) + map {S T} F := { + obj := F.app (op [0]₂) + map := fun f ↦ + { edge := F.app _ f.edge + src_eq := by rw [← FunctorToTypes.naturality, f.src_eq] + tgt_eq := by rw [← FunctorToTypes.naturality, f.tgt_eq] } + map_id := fun X ↦ OneTruncation₂.Hom.ext (by + dsimp + rw [← FunctorToTypes.naturality]) } + +@[ext] +lemma OneTruncation₂.hom_ext {S : SSet.Truncated 2} {x y : OneTruncation₂ S} {f g : x ⟶ y} : + f.edge = g.edge → f = g := OneTruncation₂.Hom.ext + +section +variable {C : Type u} [Category.{v} C] + +/-- An equivalence between the type of objects underlying a category and the type of 0-simplices in +the 2-truncated nerve. -/ +@[simps] +def OneTruncation₂.nerveEquiv : + OneTruncation₂ ((SSet.truncation 2).obj (nerve C)) ≃ C where + toFun X := X.obj' 0 + invFun X := .mk₀ X + left_inv _ := ComposableArrows.ext₀ rfl + right_inv _ := rfl + +/-- A hom equivalence over the function `OneTruncation₂.nerveEquiv`. -/ +def OneTruncation₂.nerveHomEquiv (X Y : OneTruncation₂ ((SSet.truncation 2).obj (nerve C))) : + (X ⟶ Y) ≃ (nerveEquiv X ⟶ nerveEquiv Y) where + toFun φ := eqToHom (congr_arg ComposableArrows.left φ.src_eq.symm) ≫ φ.edge.hom ≫ + eqToHom (congr_arg ComposableArrows.left φ.tgt_eq) + invFun f := + { edge := ComposableArrows.mk₁ f + src_eq := ComposableArrows.ext₀ rfl + tgt_eq := ComposableArrows.ext₀ rfl } + left_inv φ := by + ext + exact ComposableArrows.ext₁ (congr_arg ComposableArrows.left φ.src_eq).symm + (congr_arg ComposableArrows.left φ.tgt_eq).symm rfl + right_inv f := by dsimp; simp only [comp_id, id_comp]; rfl + +/-- The refl quiver underlying a nerve is isomorphic to the refl quiver underlying the category. -/ +def OneTruncation₂.ofNerve₂ (C : Type u) [Category.{u} C] : + ReflQuiv.of (OneTruncation₂ (nerveFunctor₂.obj (Cat.of C))) ≅ ReflQuiv.of C := by + apply ReflQuiv.isoOfEquiv.{u,u} OneTruncation₂.nerveEquiv OneTruncation₂.nerveHomEquiv ?_ + intro X + unfold nerveEquiv nerveHomEquiv + simp only [Cat.of_α, op_obj, ComposableArrows.obj', Fin.zero_eta, Fin.isValue, Equiv.coe_fn_mk, + nerveEquiv_apply, Nat.reduceAdd, id_edge, SimplexCategory.len_mk, id_eq, eqToHom_refl, comp_id, + id_comp, ReflQuiver.id_eq_id] + unfold nerve truncation SimplicialObject.truncation SimplexCategory.Truncated.inclusion + simp only [fullSubcategoryInclusion.obj, SimplexCategory.len_mk, Nat.reduceAdd, Fin.isValue, + SimplexCategory.toCat_map, whiskeringLeft_obj_obj, Functor.comp_map, op_obj, op_map, + Quiver.Hom.unop_op, fullSubcategoryInclusion.map, ComposableArrows.whiskerLeft_map, + Fin.zero_eta, Monotone.functor_obj, Fin.mk_one, homOfLE_leOfHom] + show X.map (𝟙 _) = _ + rw [X.map_id] + rfl + +/-- The refl quiver underlying a nerve is naturally isomorphic to the refl quiver underlying the +category. -/ +@[simps! hom_app_obj hom_app_map inv_app_obj_obj inv_app_obj_map inv_app_map] +def OneTruncation₂.ofNerve₂.natIso : + nerveFunctor₂.{u,u} ⋙ SSet.oneTruncation₂ ≅ ReflQuiv.forget := + NatIso.ofComponents (fun C => OneTruncation₂.ofNerve₂ C) (by + · intro C D F + fapply ReflPrefunctor.ext <;> simp + · exact fun _ ↦ rfl + · intro X Y f + obtain ⟨f, rfl, rfl⟩ := f + unfold SSet.oneTruncation₂ nerveFunctor₂ SSet.truncation SimplicialObject.truncation + nerveFunctor toReflPrefunctor + simp only [comp_obj, whiskeringLeft_obj_obj, ReflQuiv.of_val, Functor.comp_map, + whiskeringLeft_obj_map, whiskerLeft_app, op_obj, whiskeringRight_obj_obj, ofNerve₂, + Cat.of_α, nerveEquiv, ComposableArrows.obj', Fin.zero_eta, Fin.isValue, + ReflQuiv.comp_eq_comp, Nat.reduceAdd, SimplexCategory.len_mk, id_eq, op_map, + Quiver.Hom.unop_op, nerve_map, SimplexCategory.toCat_map, ReflPrefunctor.comp_obj, + ReflPrefunctor.comp_map] + simp [nerveHomEquiv, ReflQuiv.isoOfEquiv, ReflQuiv.isoOfQuivIso, Quiv.isoOfEquiv]) + +end + +section + +private lemma map_map_of_eq.{w} {C : Type u} [Category.{v} C] (V : Cᵒᵖ ⥤ Type w) {X Y Z : C} + {α : X ⟶ Y} {β : Y ⟶ Z} {γ : X ⟶ Z} {φ} : + α ≫ β = γ → V.map α.op (V.map β.op φ) = V.map γ.op φ := by + rintro rfl + change (V.map _ ≫ V.map _) _ = _ + rw [← map_comp]; rfl + +variable {V : SSet} + +namespace Truncated + +local notation (priority := high) "[" n "]" => SimplexCategory.mk n + +/-- The map that picks up the initial vertex of a 2-simplex, as a morphism in the 2-truncated +simplex category. -/ +def ι0₂ : [0]₂ ⟶ [2]₂ := δ₂ (n := 0) 1 ≫ δ₂ (n := 1) 1 + +/-- The map that picks up the middle vertex of a 2-simplex, as a morphism in the 2-truncated +simplex category. -/ +def ι1₂ : [0]₂ ⟶ [2]₂ := δ₂ (n := 0) 0 ≫ δ₂ (n := 1) 2 + +/-- The map that picks up the final vertex of a 2-simplex, as a morphism in the 2-truncated +simplex category. -/ +def ι2₂ : [0]₂ ⟶ [2]₂ := δ₂ (n := 0) 0 ≫ δ₂ (n := 1) 1 + +/-- The initial vertex of a 2-simplex in a 2-truncated simplicial set. -/ +def ev0₂ {V : SSet.Truncated 2} (φ : V _[2]₂) : OneTruncation₂ V := V.map ι0₂.op φ + +/-- The middle vertex of a 2-simplex in a 2-truncated simplicial set. -/ +def ev1₂ {V : SSet.Truncated 2} (φ : V _[2]₂) : OneTruncation₂ V := V.map ι1₂.op φ + +/-- The final vertex of a 2-simplex in a 2-truncated simplicial set. -/ +def ev2₂ {V : SSet.Truncated 2} (φ : V _[2]₂) : OneTruncation₂ V := V.map ι2₂.op φ + +/-- The 0th face of a 2-simplex, as a morphism in the 2-truncated simplex category. -/ +def δ0₂ : [1]₂ ⟶ [2]₂ := δ₂ (n := 1) 0 + +/-- The 1st face of a 2-simplex, as a morphism in the 2-truncated simplex category. -/ +def δ1₂ : [1]₂ ⟶ [2]₂ := δ₂ (n := 1) 1 + +/-- The 2nd face of a 2-simplex, as a morphism in the 2-truncated simplex category. -/ +def δ2₂ : [1]₂ ⟶ [2]₂ := δ₂ (n := 1) 2 + +/-- The arrow in the ReflQuiver `OneTruncation₂ V` of a 2-truncated simplicial set arising from the +0th face of a 2-simplex. -/ +def ev12₂ {V : SSet.Truncated 2} (φ : V _[2]₂) : ev1₂ φ ⟶ ev2₂ φ := + ⟨V.map δ0₂.op φ, + map_map_of_eq V (SimplexCategory.δ_comp_δ (i := 0) (j := 1) (by decide)).symm, + map_map_of_eq V rfl⟩ + +/-- The arrow in the ReflQuiver `OneTruncation₂ V` of a 2-truncated simplicial set arising from the +1st face of a 2-simplex. -/ +def ev02₂ {V : SSet.Truncated 2} (φ : V _[2]₂) : ev0₂ φ ⟶ ev2₂ φ := + ⟨V.map δ1₂.op φ, map_map_of_eq V rfl, map_map_of_eq V rfl⟩ + +/-- The arrow in the ReflQuiver `OneTruncation₂ V` of a 2-truncated simplicial set arising from the +2nd face of a 2-simplex. -/ +def ev01₂ {V : SSet.Truncated 2} (φ : V _[2]₂) : ev0₂ φ ⟶ ev1₂ φ := + ⟨V.map δ2₂.op φ, map_map_of_eq V (SimplexCategory.δ_comp_δ (j := 1) le_rfl), map_map_of_eq V rfl⟩ + + +/-- The 2-simplices in a 2-truncated simplicial set `V` generate a hom relation on the free +category on the underlying refl quiver of `V`. -/ +inductive HoRel₂ {V : SSet.Truncated 2} : + (X Y : Cat.FreeRefl (OneTruncation₂ V)) → (f g : X ⟶ Y) → Prop + | mk (φ : V _[2]₂) : + HoRel₂ _ _ + (Quot.mk _ (Quiver.Hom.toPath (ev02₂ φ))) + (Quot.mk _ ((Quiver.Hom.toPath (ev01₂ φ)).comp + (Quiver.Hom.toPath (ev12₂ φ)))) + +/-- A 2-simplex whose faces are identified with certain arrows in `OneTruncation₂ V` defines +a term of type `HoRel₂` between those arrows. -/ +theorem HoRel₂.mk' {V : SSet.Truncated 2} (φ : V _[2]₂) {X₀ X₁ X₂ : OneTruncation₂ V} + (f₀₁ : X₀ ⟶ X₁) (f₁₂ : X₁ ⟶ X₂) (f₀₂ : X₀ ⟶ X₂) + (h₀₁ : f₀₁.edge = V.map (δ₂ 2).op φ) (h₁₂ : f₁₂.edge = V.map (δ₂ 0).op φ) + (h₀₂ : f₀₂.edge = V.map (δ₂ 1).op φ) : + HoRel₂ _ _ (Quot.mk _ (Quiver.Hom.toPath f₀₂)) + (Quot.mk _ ((Quiver.Hom.toPath f₀₁).comp (Quiver.Hom.toPath f₁₂))) := by + obtain rfl : X₀ = ev0₂ φ := by + rw [← f₀₂.src_eq, h₀₂, ← FunctorToTypes.map_comp_apply, ← op_comp] + rfl + obtain rfl : X₁ = ev1₂ φ := by + rw [← f₀₁.tgt_eq, h₀₁, ← FunctorToTypes.map_comp_apply, ← op_comp] + rfl + obtain rfl : X₂ = ev2₂ φ := by + rw [← f₁₂.tgt_eq, h₁₂, ← FunctorToTypes.map_comp_apply, ← op_comp] + rfl + obtain rfl : f₀₁ = ev01₂ φ := by ext; assumption + obtain rfl : f₁₂ = ev12₂ φ := by ext; assumption + obtain rfl : f₀₂ = ev02₂ φ := by ext; assumption + constructor + +/-- The type underlying the homotopy category of a 2-truncated simplicial set `V`. -/ +def _root_.SSet.Truncated.HomotopyCategory (V : SSet.Truncated.{u} 2) : Type u := + Quotient (HoRel₂ (V := V)) + +instance (V : SSet.Truncated.{u} 2) : Category.{u} (V.HomotopyCategory) := + inferInstanceAs (Category (CategoryTheory.Quotient ..)) + +/-- A canonical functor from the free category on the refl quiver underlying a 2-truncated +simplicial set `V` to its homotopy category. -/ +def _root_.SSet.Truncated.HomotopyCategory.quotientFunctor (V : SSet.Truncated.{u} 2) : + Cat.FreeRefl (OneTruncation₂ V) ⥤ V.HomotopyCategory := + Quotient.functor _ + +/-- By `Quotient.lift_unique'` (not `Quotient.lift`) we have that `quotientFunctor V` is an +epimorphism. -/ +theorem HomotopyCategory.lift_unique' (V : SSet.Truncated.{u} 2) {D} [Category D] + (F₁ F₂ : V.HomotopyCategory ⥤ D) + (h : HomotopyCategory.quotientFunctor V ⋙ F₁ = HomotopyCategory.quotientFunctor V ⋙ F₂) : + F₁ = F₂ := + Quotient.lift_unique' (C := Cat.FreeRefl (OneTruncation₂ V)) + (HoRel₂ (V := V)) _ _ h + +/-- A map of 2-truncated simplicial sets induces a functor between homotopy categories. -/ +def mapHomotopyCategory {V W : SSet.Truncated.{u} 2} (F : V ⟶ W) : + V.HomotopyCategory ⥤ W.HomotopyCategory := + CategoryTheory.Quotient.lift _ + ((oneTruncation₂ ⋙ Cat.freeRefl).map F ⋙ HomotopyCategory.quotientFunctor W) + (by + rintro _ _ _ _ ⟨φ⟩ + apply CategoryTheory.Quotient.sound + apply HoRel₂.mk' (φ := F.app _ φ) + (f₀₁ := (oneTruncation₂.map F).map (ev01₂ φ)) + (f₀₂ := (oneTruncation₂.map F).map (ev02₂ φ)) + (f₁₂ := (oneTruncation₂.map F).map (ev12₂ φ)) + all_goals + apply FunctorToTypes.naturality) + +/-- The functor that takes a 2-truncated simplicial set to its homotopy category. -/ +def hoFunctor₂ : SSet.Truncated.{u} 2 ⥤ Cat.{u,u} where + obj V := Cat.of (V.HomotopyCategory) + map {S T} F := mapHomotopyCategory F + map_id S := by + apply Quotient.lift_unique' + simp [mapHomotopyCategory, Quotient.lift_spec] + exact Eq.trans (Functor.id_comp ..) (Functor.comp_id _).symm + map_comp {S T U} F G := by + apply Quotient.lift_unique' + simp [mapHomotopyCategory, SSet.Truncated.HomotopyCategory.quotientFunctor] + rw [Quotient.lift_spec, Cat.comp_eq_comp, Cat.comp_eq_comp, ← Functor.assoc, Functor.assoc, + Quotient.lift_spec, Functor.assoc, Quotient.lift_spec] + +theorem hoFunctor₂_naturality {X Y : SSet.Truncated.{u} 2} (f : X ⟶ Y) : + (oneTruncation₂ ⋙ Cat.freeRefl).map f ⋙ SSet.Truncated.HomotopyCategory.quotientFunctor Y = + SSet.Truncated.HomotopyCategory.quotientFunctor X ⋙ mapHomotopyCategory f := rfl + +end Truncated + +/-- The functor that takes a simplicial set to its homotopy category by passing through the +2-truncation. -/ +def hoFunctor : SSet.{u} ⥤ Cat.{u, u} := SSet.truncation 2 ⋙ Truncated.hoFunctor₂ + +end + +end + +end SSet diff --git a/Mathlib/CategoryTheory/Category/Quiv.lean b/Mathlib/CategoryTheory/Category/Quiv.lean index b8ad841e24d51..eeacc38f92cd2 100644 --- a/Mathlib/CategoryTheory/Category/Quiv.lean +++ b/Mathlib/CategoryTheory/Category/Quiv.lean @@ -82,6 +82,71 @@ end Cat namespace Quiv +section +variable {V W : Quiv} (e : V ≅ W) + +/-- An isomorphism of quivers defines an equivalence on carrier types. -/ +@[simps] +def equivOfIso : V ≃ W where + toFun := e.hom.obj + invFun := e.inv.obj + left_inv := Prefunctor.congr_obj e.hom_inv_id + right_inv := Prefunctor.congr_obj e.inv_hom_id + +@[simp] +lemma inv_obj_hom_obj_of_iso (X : V) : e.inv.obj (e.hom.obj X) = X := (equivOfIso e).left_inv X + +@[simp] +lemma hom_obj_inv_obj_of_iso (Y : W) : e.hom.obj (e.inv.obj Y) = Y := (equivOfIso e).right_inv Y + +lemma hom_map_inv_map_of_iso {V W : Quiv} (e : V ≅ W) {X Y : W} (f : X ⟶ Y) : + e.hom.map (e.inv.map f) = Quiver.homOfEq f (by simp) (by simp) := by + rw [← Prefunctor.comp_map] + exact (Prefunctor.congr_hom e.inv_hom_id.symm f).symm + +lemma inv_map_hom_map_of_iso {V W : Quiv} (e : V ≅ W) {X Y : V} (f : X ⟶ Y) : + e.inv.map (e.hom.map f) = Quiver.homOfEq f (by simp) (by simp) := + hom_map_inv_map_of_iso e.symm f + +/-- An isomorphism of quivers defines an equivalence on hom types. -/ +@[simps] +def homEquivOfIso {V W : Quiv} (e : V ≅ W) {X Y : V} : + (X ⟶ Y) ≃ (e.hom.obj X ⟶ e.hom.obj Y) where + toFun f := e.hom.map f + invFun g := Quiver.homOfEq (e.inv.map g) (by simp) (by simp) + left_inv f := by simp [inv_map_hom_map_of_iso] + right_inv g := by simp [hom_map_inv_map_of_iso] + +end + +section +variable {V W : Type u } [Quiver V] [Quiver W] + (e : V ≃ W) (he : ∀ X Y : V, (X ⟶ Y) ≃ (e X ⟶ e Y)) + +include he in +@[simp] +lemma homOfEq_map_homOfEq {X Y : V} (f : X ⟶ Y) {X' Y' : V} (hX : X = X') (hY : Y = Y') + {X'' Y'' : W} (hX' : e X' = X'') (hY' : e Y' = Y'') : + Quiver.homOfEq (he _ _ (Quiver.homOfEq f hX hY)) hX' hY' = + Quiver.homOfEq (he _ _ f) (by rw [hX, hX']) (by rw [hY, hY']) := by + subst hX hY hX' hY' + rfl + +/-- Compatible equivalences of types and hom-types induce an isomorphism of quivers. -/ +def isoOfEquiv : Quiv.of V ≅ Quiv.of W where + hom := Prefunctor.mk e (he _ _) + inv := + { obj := e.symm + map {X Y} f := (he _ _).symm (Quiver.homOfEq f (by simp) (by simp)) } + hom_inv_id := Prefunctor.ext' e.left_inv (fun X Y f ↦ by + dsimp [Quiv.id_eq_id, Quiv.comp_eq_comp] + apply (he _ _).injective + apply Quiver.homOfEq_injective (X' := e X) (Y' := e Y) (by simp) (by simp) + simp) + inv_hom_id := Prefunctor.ext' e.right_inv (by simp [Quiv.id_eq_id, Quiv.comp_eq_comp]) + +end + /-- Any prefunctor into a category lifts to a functor from the path category. -/ @[simps] def lift {V : Type u} [Quiver.{v + 1} V] {C : Type*} [Category C] (F : Prefunctor V C) : diff --git a/Mathlib/CategoryTheory/Category/ReflQuiv.lean b/Mathlib/CategoryTheory/Category/ReflQuiv.lean index f26afec53930b..8ec09b21eeb6e 100644 --- a/Mathlib/CategoryTheory/Category/ReflQuiv.lean +++ b/Mathlib/CategoryTheory/Category/ReflQuiv.lean @@ -15,7 +15,7 @@ The category `ReflQuiv` of (bundled) reflexive quivers, and the free/forgetful -/ namespace CategoryTheory -universe v u +universe v u v₁ v₂ u₁ u₂ /-- Category of refl quivers. -/ @[nolint checkUnivs] @@ -70,11 +70,33 @@ 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 +instance forgetToQuiv.Faithful : Functor.Faithful forgetToQuiv where map_injective := fun hyp ↦ forgetToQuiv_faithful _ _ hyp theorem forget_forgetToQuiv : forget ⋙ forgetToQuiv = Quiv.forget := rfl +/-- An isomorphism of quivers lifts to an isomorphism of reflexive quivers given a suitable +compatibility with the identities. -/ +def isoOfQuivIso {V W : Type u} [ReflQuiver V] [ReflQuiver W] + (e : Quiv.of V ≅ Quiv.of W) + (h_id : ∀ (X : V), e.hom.map (𝟙rq X) = ReflQuiver.id (obj := W) (e.hom.obj X)) : + ReflQuiv.of V ≅ ReflQuiv.of W where + hom := ReflPrefunctor.mk e.hom h_id + inv := ReflPrefunctor.mk e.inv + (fun Y => (Quiv.homEquivOfIso e).injective (by simp [Quiv.hom_map_inv_map_of_iso, h_id])) + hom_inv_id := by + apply forgetToQuiv.map_injective + exact e.hom_inv_id + inv_hom_id := by + apply forgetToQuiv.map_injective + exact e.inv_hom_id + +/-- Compatible equivalences of types and hom-types induce an isomorphism of reflexive quivers. -/ +def isoOfEquiv {V W : Type u } [ReflQuiver V] [ReflQuiver W] (e : V ≃ W) + (he : ∀ (X Y : V), (X ⟶ Y) ≃ (e X ⟶ e Y)) + (h_id : ∀ (X : V), he _ _ (𝟙rq X) = ReflQuiver.id (obj := W) (e X)) : + ReflQuiv.of V ≅ ReflQuiv.of W := isoOfQuivIso (Quiv.isoOfEquiv e he) h_id + end ReflQuiv namespace ReflPrefunctor diff --git a/Mathlib/CategoryTheory/Quotient.lean b/Mathlib/CategoryTheory/Quotient.lean index 612ea6e8c5677..6ba3fc3efd70f 100644 --- a/Mathlib/CategoryTheory/Quotient.lean +++ b/Mathlib/CategoryTheory/Quotient.lean @@ -28,6 +28,20 @@ instance (C) [Quiver C] : Inhabited (HomRel C) where namespace CategoryTheory +section + +variable {C D : Type*} [Category C] [Category D] (F : C ⥤ D) + +/-- A functor induces a `HomRel` on its domain, relating those maps that have the same image. -/ +def Functor.homRel : HomRel C := + fun _ _ f g ↦ F.map f = F.map g + +@[simp] +lemma Functor.homRel_iff {X Y : C} (f g : X ⟶ Y) : + F.homRel f g ↔ F.map f = F.map g := Iff.rfl + +end + variable {C : Type _} [Category C] (r : HomRel C) /-- A `HomRel` is a congruence when it's an equivalence on every hom-set, and it can be composed @@ -40,6 +54,16 @@ class Congruence : Prop where /-- Postcomposition with an arrow respects `r`. -/ compRight : ∀ {X Y Z} {f f' : X ⟶ Y} (g : Y ⟶ Z), r f f' → r (f ≫ g) (f' ≫ g) +/-- For `F : C ⥤ D`, `F.homRel` is a congruence.-/ +instance Functor.congruence_homRel {C D : Type*} [Category C] [Category D] (F : C ⥤ D) : + Congruence F.homRel where + equivalence := + { refl := fun _ ↦ rfl + symm := by aesop + trans := by aesop } + compLeft := by aesop + compRight := by aesop + /-- A type synonym for `C`, thought of as the objects of the quotient category. -/ @[ext] structure Quotient (r : HomRel C) where @@ -140,6 +164,16 @@ theorem functor_map_eq_iff [h : Congruence r] {X Y : C} (f f' : X ⟶ Y) : rw [Equivalence.quot_mk_eq_iff, compClosure_eq_self r] simpa only [compClosure_eq_self r] using h.equivalence +theorem functor_homRel_eq_compClosure_eqvGen {X Y : C} (f g : X ⟶ Y) : + (functor r).homRel f g ↔ Relation.EqvGen (@CompClosure C _ r X Y) f g := + Quot.eq + +theorem compClosure.congruence : + Congruence fun X Y => Relation.EqvGen (@CompClosure C _ r X Y) := by + convert inferInstanceAs (Congruence (functor r).homRel) + ext + rw [functor_homRel_eq_compClosure_eqvGen] + variable {D : Type _} [Category D] (F : C ⥤ D) /-- The induced functor on the quotient category. -/ diff --git a/Mathlib/Combinatorics/Quiver/Basic.lean b/Mathlib/Combinatorics/Quiver/Basic.lean index 4b5b56f23e783..ed4441ac1f8be 100644 --- a/Mathlib/Combinatorics/Quiver/Basic.lean +++ b/Mathlib/Combinatorics/Quiver/Basic.lean @@ -81,4 +81,32 @@ theorem empty_arrow {V : Type u} (a b : Empty V) : (a ⟶ b) = PEmpty := rfl /-- A quiver is thin if it has no parallel arrows. -/ abbrev IsThin (V : Type u) [Quiver V] : Prop := ∀ a b : V, Subsingleton (a ⟶ b) + +section + +variable {V : Type*} [Quiver V] + +/-- An arrow in a quiver can be transported across equalities between the source and target +objects. -/ +def homOfEq {X Y : V} (f : X ⟶ Y) {X' Y' : V} (hX : X = X') (hY : Y = Y') : X' ⟶ Y' := by + subst hX hY + exact f + +@[simp] +lemma homOfEq_trans {X Y : V} (f : X ⟶ Y) {X' Y' : V} (hX : X = X') (hY : Y = Y') + {X'' Y'' : V} (hX' : X' = X'') (hY' : Y' = Y'') : + homOfEq (homOfEq f hX hY) hX' hY' = homOfEq f (hX.trans hX') (hY.trans hY') := by + subst hX hY hX' hY' + rfl + +lemma homOfEq_injective {X X' Y Y' : V} (hX : X = X') (hY : Y = Y') + {f g : X ⟶ Y} (h : Quiver.homOfEq f hX hY = Quiver.homOfEq g hX hY) : f = g := by + subst hX hY + exact h + +@[simp] +lemma homOfEq_rfl {X Y : V} (f : X ⟶ Y) : Quiver.homOfEq f rfl rfl = f := rfl + +end + end Quiver diff --git a/Mathlib/Combinatorics/Quiver/Prefunctor.lean b/Mathlib/Combinatorics/Quiver/Prefunctor.lean index 98020fcc9bf1b..2a0468c73b08f 100644 --- a/Mathlib/Combinatorics/Quiver/Prefunctor.lean +++ b/Mathlib/Combinatorics/Quiver/Prefunctor.lean @@ -45,6 +45,18 @@ theorem ext {V : Type u} [Quiver.{v₁} V] {W : Type u₂} [Quiver.{v₂} W] {F funext X Y f simpa using h_map X Y f +/-- This may be a more useful form of `Prefunctor.ext`. -/ +theorem ext' {V W : Type u} [Quiver V] [Quiver W] {F G : Prefunctor V W} + (h_obj : ∀ X, F.obj X = G.obj X) + (h_map : ∀ (X Y : V) (f : X ⟶ Y), + F.map f = Quiver.homOfEq (G.map f) (h_obj _).symm (h_obj _).symm) : F = G := by + obtain ⟨Fobj, Fmap⟩ := F + obtain ⟨Gobj, Gmap⟩ := G + obtain rfl : Fobj = Gobj := funext h_obj + simp only [mk.injEq, heq_eq_eq, true_and] + ext X Y f + simpa only [Quiver.homOfEq_rfl] using h_map X Y f + /-- The identity morphism between quivers. -/ @[simps] def id (V : Type*) [Quiver V] : Prefunctor V V where @@ -88,4 +100,21 @@ theorem congr_map {U V : Type*} [Quiver U] [Quiver V] (F : U ⥤q V) {X Y : U} { (h : f = g) : F.map f = F.map g := by rw [h] +/-- An equality of prefunctors gives an equality on objects. -/ +theorem congr_obj {U V : Type*} [Quiver U] [Quiver V] {F G : U ⥤q V} (e : F = G) (X : U) : + F.obj X = G.obj X := by cases e; rfl + +/-- An equality of prefunctors gives an equality on homs. -/ +theorem congr_hom {U V : Type*} [Quiver U] [Quiver V] {F G : U ⥤q V} (e : F = G) {X Y : U} + (f : X ⟶ Y) : Quiver.homOfEq (F.map f) (congr_obj e X) (congr_obj e Y) = G.map f := by + subst e + simp + +/-- Prefunctors commute with `homOfEq`. -/ +@[simp] +theorem homOfEq_map {U V : Type*} [Quiver U] [Quiver V] (F : U ⥤q V) {X Y : U} (f : X ⟶ Y) + {X' Y' : U} (hX : X = X') (hY : Y = Y') : + F.map (Quiver.homOfEq f hX hY) = + Quiver.homOfEq (F.map f) (congr_arg F.obj hX) (congr_arg F.obj hY) := by subst hX hY; simp + end Prefunctor diff --git a/Mathlib/Combinatorics/Quiver/ReflQuiver.lean b/Mathlib/Combinatorics/Quiver/ReflQuiver.lean index b0f9c85ac7250..33a3689756282 100644 --- a/Mathlib/Combinatorics/Quiver/ReflQuiver.lean +++ b/Mathlib/Combinatorics/Quiver/ReflQuiver.lean @@ -31,6 +31,10 @@ class ReflQuiver (obj : Type u) extends Quiver.{v} obj : Type max u v where /-- Notation for the identity morphism in a category. -/ scoped notation "𝟙rq" => ReflQuiver.id -- type as \b1 +@[simp] +theorem ReflQuiver.homOfEq_id {V : Type*} [ReflQuiver V] {X X' : V} (hX : X = X') : + Quiver.homOfEq (𝟙rq X) hX hX = 𝟙rq X' := by subst hX ; rfl + instance catToReflQuiver {C : Type u} [inst : Category.{v} C] : ReflQuiver.{v+1, u} C := { inst with } @@ -66,6 +70,18 @@ theorem ext {V : Type u} [ReflQuiver.{v₁} V] {W : Type u₂} [ReflQuiver.{v₂ funext X Y f simpa using h_map X Y f +/-- This may be a more useful form of `ReflPrefunctor.ext`. -/ +theorem ext' {V W : Type u} [ReflQuiver.{v} V] [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 = Quiver.homOfEq (G.map f) (h_obj _).symm (h_obj _).symm) : F = G := by + obtain ⟨Fpre, Fid⟩ := F + obtain ⟨Gpre, Gid⟩ := G + simp at h_obj h_map + obtain rfl : Fpre = Gpre := Prefunctor.ext' (V := V) (W := W) h_obj h_map + rfl + /-- The identity morphism between reflexive quivers. -/ @[simps!] def id (V : Type*) [ReflQuiver V] : ReflPrefunctor V V where From 2e75934cfaa46033c497006007630344dbafeac4 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Fri, 20 Dec 2024 03:50:08 +0000 Subject: [PATCH 795/829] feat(RingTheory/Invariant): Existence of Frobenius elements (#19926) This PR proves the existence of Frobenius elements in algebraic number theory. --- Mathlib/RingTheory/Invariant.lean | 233 +++++++++++++++++++++++++++++- 1 file changed, 225 insertions(+), 8 deletions(-) diff --git a/Mathlib/RingTheory/Invariant.lean b/Mathlib/RingTheory/Invariant.lean index 46ac790026076..075f78efeb749 100644 --- a/Mathlib/RingTheory/Invariant.lean +++ b/Mathlib/RingTheory/Invariant.lean @@ -3,23 +3,36 @@ Copyright (c) 2024 Thomas Browning. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning -/ -import Mathlib.Algebra.Polynomial.GroupRingAction -import Mathlib.Algebra.Polynomial.Lifts -import Mathlib.RingTheory.IntegralClosure.Algebra.Defs +import Mathlib.FieldTheory.Fixed +import Mathlib.RingTheory.Ideal.Over /-! -# Invariant Extensions of Rings +# Invariant Extensions of Rings and Frobenius Elements + +In algebraic number theory, if `L/K` is a finite Galois extension of number fields, with rings of +integers `𝓞L/𝓞K`, and if `q` is prime ideal of `𝓞L` lying over a prime ideal `p` of `𝓞K`, then +there exists a **Frobenius element** `Frob p` in `Gal(L/K)` with the property that +`Frob p x ≡ x ^ #(𝓞K/p) (mod q)` for all `x ∈ 𝓞L`. + +This file proves the existence of Frobenius elements in a more general setting. Given an extension of rings `B/A` and an action of `G` on `B`, we introduce a predicate `Algebra.IsInvariant A B G` which states that every fixed point of `B` lies in the image of `A`. ## Main statements -* `Algebra.IsInvariant.isIntegral`: If `G` is finite, then `B/A` is an integral extension. -TODO: Prove the existence of Frobenius elements in this general setting (PR #17717). +Let `G` be a finite group acting on a commutative ring `B` satisfying `Algebra.IsInvariant A B G`. + +* `Algebra.IsInvariant.isIntegral`: `B/A` is an integral extension. +* `Algebra.IsInvariant.exists_smul_of_under_eq`: `G` acts transitivity on the prime ideals of `B` + lying above a given prime ideal of `A`. +* `IsFractionRing.stabilizerHom_surjective`: if `Q` is a prime ideal of `B` lying over a prime + ideal `P` of `A`, then the stabilizer subgroup of `Q` surjects onto `Aut(Frac(B/Q)/Frac(A/P))`. -/ +open scoped Pointwise + namespace Algebra variable (A B G : Type*) [CommSemiring A] [Semiring B] [Algebra A B] @@ -33,7 +46,7 @@ by `G` is `smul_algebraMap` (assuming `SMulCommClass A B G`). -/ end Algebra -section IsIntegral +section transitivity variable (A B G : Type*) [CommRing A] [CommRing B] [Algebra A B] [Group G] [MulSemiringAction G B] @@ -87,6 +100,210 @@ theorem isIntegral [Finite G] : Algebra.IsIntegral A B := by (charpoly_mem_lifts A B G b) (monic_charpoly G b) exact ⟨p, hp2, by rw [← eval_map, hp1, eval_charpoly]⟩ +/-- `G` acts transitively on the prime ideals of `B` above a given prime ideal of `A`. -/ +theorem exists_smul_of_under_eq [Finite G] [SMulCommClass G A B] + (P Q : Ideal B) [hP : P.IsPrime] [hQ : Q.IsPrime] + (hPQ : P.under A = Q.under A) : + ∃ g : G, Q = g • P := by + cases nonempty_fintype G + have : ∀ (P Q : Ideal B) [P.IsPrime] [Q.IsPrime], P.under A = Q.under A → + ∃ g ∈ (⊤ : Finset G), Q ≤ g • P := by + intro P Q hP hQ hPQ + rw [← Ideal.subset_union_prime 1 1 (fun _ _ _ _ ↦ hP.smul _)] + intro b hb + suffices h : ∃ g ∈ Finset.univ, g • b ∈ P by + obtain ⟨g, -, hg⟩ := h + apply Set.mem_biUnion (Finset.mem_univ g⁻¹) (Ideal.mem_inv_pointwise_smul_iff.mpr hg) + obtain ⟨a, ha⟩ := isInvariant (A := A) (∏ g : G, g • b) (Finset.smul_prod_perm b) + rw [← hP.prod_mem_iff, ← ha, ← P.mem_comap, ← P.under_def A, + hPQ, Q.mem_comap, ha, hQ.prod_mem_iff] + exact ⟨1, Finset.mem_univ 1, (one_smul G b).symm ▸ hb⟩ + obtain ⟨g, -, hg⟩ := this P Q hPQ + obtain ⟨g', -, hg'⟩ := this Q (g • P) ((P.under_smul A g).trans hPQ).symm + exact ⟨g, le_antisymm hg (smul_eq_of_le_smul (hg.trans hg') ▸ hg')⟩ + end Algebra.IsInvariant -end IsIntegral +end transitivity + +section surjectivity + +open IsScalarTower NoZeroSMulDivisors Polynomial + +variable {A B : Type*} [CommRing A] [CommRing B] [Algebra A B] + (G : Type*) [Group G] [Finite G] [MulSemiringAction G B] [SMulCommClass G A B] + (P : Ideal A) (Q : Ideal B) [Q.IsPrime] [Q.LiesOver P] + variable (K L : Type*) [Field K] [Field L] + [Algebra (A ⧸ P) K] [Algebra (B ⧸ Q) L] + [Algebra (A ⧸ P) L] [IsScalarTower (A ⧸ P) (B ⧸ Q) L] + [Algebra K L] [IsScalarTower (A ⧸ P) K L] + [Algebra.IsInvariant A B G] + +/-- A technical lemma for `fixed_of_fixed1`. -/ +private theorem fixed_of_fixed1_aux1 [DecidableEq (Ideal B)] : + ∃ a b : B, (∀ g : G, g • a = a) ∧ a ∉ Q ∧ + ∀ g : G, algebraMap B (B ⧸ Q) (g • b) = algebraMap B (B ⧸ Q) (if g • Q = Q then a else 0) := by + obtain ⟨_⟩ := nonempty_fintype G + let P := ((Finset.univ : Finset G).filter (fun g ↦ g • Q ≠ Q)).inf (fun g ↦ g • Q) + have h1 : ¬ P ≤ Q := by + rw [Ideal.IsPrime.inf_le' inferInstance] + rintro ⟨g, hg1, hg2⟩ + exact (Finset.mem_filter.mp hg1).2 (smul_eq_of_smul_le hg2) + obtain ⟨b, hbP, hbQ⟩ := SetLike.not_le_iff_exists.mp h1 + replace hbP : ∀ g : G, g • Q ≠ Q → b ∈ g • Q := + fun g hg ↦ (Finset.inf_le (Finset.mem_filter.mpr ⟨Finset.mem_univ g, hg⟩) : P ≤ g • Q) hbP + let f := MulSemiringAction.charpoly G b + obtain ⟨q, hq, hq0⟩ := + (f.map (algebraMap B (B ⧸ Q))).exists_eq_pow_rootMultiplicity_mul_and_not_dvd + (Polynomial.map_monic_ne_zero (MulSemiringAction.monic_charpoly G b)) 0 + rw [map_zero, sub_zero] at hq hq0 + let j := (f.map (algebraMap B (B ⧸ Q))).rootMultiplicity 0 + let k := q.natDegree + let r := ∑ i ∈ Finset.range (k + 1), Polynomial.monomial i (f.coeff (i + j)) + have hr : r.map (algebraMap B (B ⧸ Q)) = q := by + ext n + rw [Polynomial.coeff_map, Polynomial.finset_sum_coeff] + simp only [Polynomial.coeff_monomial, Finset.sum_ite_eq', Finset.mem_range_succ_iff] + split_ifs with hn + · rw [← Polynomial.coeff_map, hq, Polynomial.coeff_X_pow_mul] + · rw [map_zero, eq_comm, Polynomial.coeff_eq_zero_of_natDegree_lt (lt_of_not_le hn)] + have hf : f.eval b = 0 := MulSemiringAction.eval_charpoly G b + have hr : r.eval b ∈ Q := by + rw [← Ideal.Quotient.eq_zero_iff_mem, ← Ideal.Quotient.algebraMap_eq] at hbQ ⊢ + replace hf := congrArg (algebraMap B (B ⧸ Q)) hf + rw [← Polynomial.eval₂_at_apply, ← Polynomial.eval_map] at hf ⊢ + rwa [map_zero, hq, ← hr, Polynomial.eval_mul, Polynomial.eval_pow, Polynomial.eval_X, + mul_eq_zero, or_iff_right (pow_ne_zero _ hbQ)] at hf + let a := f.coeff j + have ha : ∀ g : G, g • a = a := MulSemiringAction.smul_coeff_charpoly b j + have hr' : ∀ g : G, g • Q ≠ Q → a - r.eval b ∈ g • Q := by + intro g hg + have hr : r = ∑ i ∈ Finset.range (k + 1), Polynomial.monomial i (f.coeff (i + j)) := rfl + rw [← Ideal.neg_mem_iff, neg_sub, hr, Finset.sum_range_succ', Polynomial.eval_add, + Polynomial.eval_monomial, zero_add, pow_zero, mul_one, add_sub_cancel_right] + simp only [ ← Polynomial.monomial_mul_X] + rw [← Finset.sum_mul, Polynomial.eval_mul_X] + exact Ideal.mul_mem_left (g • Q) _ (hbP g hg) + refine ⟨a, a - r.eval b, ha, ?_, fun h ↦ ?_⟩ + · rwa [← Ideal.Quotient.eq_zero_iff_mem, ← Ideal.Quotient.algebraMap_eq, ← Polynomial.coeff_map, + ← zero_add j, hq, Polynomial.coeff_X_pow_mul, ← Polynomial.X_dvd_iff] + · rw [← sub_eq_zero, ← map_sub, Ideal.Quotient.algebraMap_eq, Ideal.Quotient.eq_zero_iff_mem, + ← Ideal.smul_mem_pointwise_smul_iff (a := h⁻¹), smul_sub, inv_smul_smul] + simp only [← eq_inv_smul_iff (g := h), eq_comm (a := Q)] + split_ifs with hh + · rwa [ha, sub_sub_cancel_left, hh, Q.neg_mem_iff] + · rw [smul_zero, sub_zero] + exact hr' h⁻¹ hh + +/-- A technical lemma for `fixed_of_fixed1`. -/ +private theorem fixed_of_fixed1_aux2 [DecidableEq (Ideal B)] (b₀ : B) + (hx : ∀ g : G, g • Q = Q → algebraMap B (B ⧸ Q) (g • b₀) = algebraMap B (B ⧸ Q) b₀) : + ∃ a b : B, (∀ g : G, g • a = a) ∧ a ∉ Q ∧ + (∀ g : G, algebraMap B (B ⧸ Q) (g • b) = + algebraMap B (B ⧸ Q) (if g • Q = Q then a * b₀ else 0)) := by + obtain ⟨a, b, ha1, ha2, hb⟩ := fixed_of_fixed1_aux1 G Q + refine ⟨a, b * b₀, ha1, ha2, fun g ↦ ?_⟩ + rw [smul_mul', map_mul, hb] + specialize hb g + split_ifs with hg + · rw [map_mul, hx g hg] + · rw [map_zero, zero_mul] + +/-- A technical lemma for `fixed_of_fixed1`. -/ +private theorem fixed_of_fixed1_aux3 [NoZeroDivisors B] {b : B} {i j : ℕ} {p : Polynomial A} + (h : p.map (algebraMap A B) = (X - C b) ^ i * X ^ j) (f : B ≃ₐ[A] B) (hi : i ≠ 0) : + f b = b := by + by_cases ha : b = 0 + · rw [ha, map_zero] + have hf := congrArg (eval b) (congrArg (Polynomial.mapAlgHom f.toAlgHom) h) + rw [coe_mapAlgHom, map_map, f.toAlgHom.comp_algebraMap, h] at hf + simp_rw [Polynomial.map_mul, Polynomial.map_pow, Polynomial.map_sub, map_X, map_C, + eval_mul, eval_pow, eval_sub, eval_X, eval_C, sub_self, zero_pow hi, zero_mul, + zero_eq_mul, or_iff_left (pow_ne_zero j ha), pow_eq_zero_iff hi, sub_eq_zero] at hf + exact hf.symm + +/-- This theorem will be made redundant by `IsFractionRing.stabilizerHom_surjective`. -/ +private theorem fixed_of_fixed1 [NoZeroSMulDivisors (B ⧸ Q) L] (f : L ≃ₐ[K] L) (b : B ⧸ Q) + (hx : ∀ g : MulAction.stabilizer G Q, Ideal.Quotient.stabilizerHom Q P G g b = b) : + f (algebraMap (B ⧸ Q) L b) = (algebraMap (B ⧸ Q) L b) := by + classical + cases nonempty_fintype G + obtain ⟨b₀, rfl⟩ := Ideal.Quotient.mk_surjective b + rw [← Ideal.Quotient.algebraMap_eq] + obtain ⟨a, b, ha1, ha2, hb⟩ := fixed_of_fixed1_aux2 G Q b₀ (fun g hg ↦ hx ⟨g, hg⟩) + obtain ⟨M, key⟩ := (mem_lifts _).mp (Algebra.IsInvariant.charpoly_mem_lifts A B G b) + replace key := congrArg (map (algebraMap B (B ⧸ Q))) key + rw [map_map, ← algebraMap_eq, algebraMap_eq A (A ⧸ P) (B ⧸ Q), + ← map_map, MulSemiringAction.charpoly, Polynomial.map_prod] at key + have key₀ : ∀ g : G, (X - C (g • b)).map (algebraMap B (B ⧸ Q)) = + if g • Q = Q then X - C (algebraMap B (B ⧸ Q) (a * b₀)) else X := by + intro g + rw [Polynomial.map_sub, map_X, map_C, hb] + split_ifs + · rfl + · rw [map_zero, map_zero, sub_zero] + simp only [key₀, Finset.prod_ite, Finset.prod_const] at key + replace key := congrArg (map (algebraMap (B ⧸ Q) L)) key + rw [map_map, ← algebraMap_eq, algebraMap_eq (A ⧸ P) K L, + ← map_map, Polynomial.map_mul, Polynomial.map_pow, Polynomial.map_pow, Polynomial.map_sub, + map_X, map_C] at key + replace key := fixed_of_fixed1_aux3 key f (Finset.card_ne_zero_of_mem + (Finset.mem_filter.mpr ⟨Finset.mem_univ 1, one_smul G Q⟩)) + simp only [map_mul] at key + obtain ⟨a, rfl⟩ := Algebra.IsInvariant.isInvariant (A := A) a ha1 + rwa [← algebraMap_apply A B (B ⧸ Q), algebraMap_apply A (A ⧸ P) (B ⧸ Q), + ← algebraMap_apply, algebraMap_apply (A ⧸ P) K L, f.commutes, mul_right_inj'] at key + rwa [← algebraMap_apply, algebraMap_apply (A ⧸ P) (B ⧸ Q) L, + ← algebraMap_apply A (A ⧸ P) (B ⧸ Q), algebraMap_apply A B (B ⧸ Q), + Ne, algebraMap_eq_zero_iff, Ideal.Quotient.algebraMap_eq, Ideal.Quotient.eq_zero_iff_mem] + +variable [IsFractionRing (A ⧸ P) K] [IsFractionRing (B ⧸ Q) L] + +/-- If `Q` lies over `P`, then the stabilizer of `Q` acts on `Frac(B/Q)/Frac(A/P)`. -/ +noncomputable def IsFractionRing.stabilizerHom : MulAction.stabilizer G Q →* (L ≃ₐ[K] L) := + MonoidHom.comp (IsFractionRing.fieldEquivOfAlgEquivHom K L) (Ideal.Quotient.stabilizerHom Q P G) + +/-- This theorem will be made redundant by `IsFractionRing.stabilizerHom_surjective`. -/ +private theorem fixed_of_fixed2 (f : L ≃ₐ[K] L) (x : L) + (hx : ∀ g : MulAction.stabilizer G Q, IsFractionRing.stabilizerHom G P Q K L g x = x) : + f x = x := by + obtain ⟨_⟩ := nonempty_fintype G + have : P.IsPrime := Ideal.over_def Q P ▸ Ideal.IsPrime.under A Q + have : Algebra.IsIntegral A B := Algebra.IsInvariant.isIntegral A B G + obtain ⟨x, y, hy, rfl⟩ := IsFractionRing.div_surjective (A := B ⧸ Q) x + obtain ⟨b, a, ha, h⟩ := (Algebra.IsAlgebraic.isAlgebraic (R := A ⧸ P) y).exists_smul_eq_mul x hy + replace ha : algebraMap (A ⧸ P) L a ≠ 0 := by + rwa [Ne, algebraMap_apply (A ⧸ P) K L, algebraMap_eq_zero_iff, algebraMap_eq_zero_iff] + replace hy : algebraMap (B ⧸ Q) L y ≠ 0 := + mt (algebraMap_eq_zero_iff (B ⧸ Q) L).mp (nonZeroDivisors.ne_zero hy) + replace h : algebraMap (B ⧸ Q) L x / algebraMap (B ⧸ Q) L y = + algebraMap (B ⧸ Q) L b / algebraMap (A ⧸ P) L a := by + rw [mul_comm, Algebra.smul_def, mul_comm] at h + rw [div_eq_div_iff hy ha, ← map_mul, ← h, map_mul, ← algebraMap_apply] + simp only [h, map_div₀, algebraMap_apply (A ⧸ P) K L, AlgEquiv.commutes] at hx ⊢ + simp only [← algebraMap_apply, div_left_inj' ha] at hx ⊢ + exact fixed_of_fixed1 G P Q K L f b (fun g ↦ IsFractionRing.injective (B ⧸ Q) L + ((IsFractionRing.fieldEquivOfAlgEquiv_algebraMap K L L + (Ideal.Quotient.stabilizerHom Q P G g) b).symm.trans (hx g))) + +/-- The stabilizer subgroup of `Q` surjects onto `Aut(Frac(B/Q)/Frac(A/P))`. -/ +theorem IsFractionRing.stabilizerHom_surjective : + Function.Surjective (stabilizerHom G P Q K L) := by + let _ := MulSemiringAction.compHom L (stabilizerHom G P Q K L) + intro f + obtain ⟨g, hg⟩ := FixedPoints.toAlgAut_surjective (MulAction.stabilizer G Q) L + (AlgEquiv.ofRingEquiv (f := f) (fun x ↦ fixed_of_fixed2 G P Q K L f x x.2)) + exact ⟨g, by rwa [AlgEquiv.ext_iff] at hg ⊢⟩ + +/-- The stabilizer subgroup of `Q` surjects onto `Aut((B/Q)/(A/P))`. -/ +theorem Ideal.Quotient.stabilizerHom_surjective : + Function.Surjective (Ideal.Quotient.stabilizerHom Q P G) := by + have : P.IsPrime := Ideal.over_def Q P ▸ Ideal.IsPrime.under A Q + let _ := FractionRing.liftAlgebra (A ⧸ P) (FractionRing (B ⧸ Q)) + have key := IsFractionRing.stabilizerHom_surjective G P Q + (FractionRing (A ⧸ P)) (FractionRing (B ⧸ Q)) + rw [IsFractionRing.stabilizerHom, MonoidHom.coe_comp] at key + exact key.of_comp_left (IsFractionRing.fieldEquivOfAlgEquivHom_injective (A ⧸ P) (B ⧸ Q) + (FractionRing (A ⧸ P)) (FractionRing (B ⧸ Q))) + +end surjectivity From 725fd7ac1c92cf4d7984b25646052f5a6e717d31 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 20 Dec 2024 03:50:09 +0000 Subject: [PATCH 796/829] chore: backport fixes for leanprover/lean4#6397 (#20097) These fixes are all of the form "a terminal simp only which was unnecessarily fragile", i.e., can be replaced by a `simp` that works both on `v4.15.0-rc1` and under leanprover/lean4#6397. --- Mathlib/Algebra/BigOperators/Group/List.lean | 2 +- Mathlib/Analysis/Convex/StoneSeparation.lean | 3 +-- Mathlib/Computability/TuringMachine.lean | 3 +-- Mathlib/GroupTheory/PushoutI.lean | 4 +--- Mathlib/ModelTheory/Algebra/Ring/Definability.lean | 6 +----- Mathlib/ModelTheory/Syntax.lean | 2 +- Mathlib/Testing/Plausible/Functions.lean | 3 +-- 7 files changed, 7 insertions(+), 16 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index 0eb24a9825d47..5641f81980c14 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -649,7 +649,7 @@ lemma length_sigma {σ : α → Type*} (l₁ : List α) (l₂ : ∀ a, List (σ lemma ranges_flatten : ∀ (l : List ℕ), l.ranges.flatten = range l.sum | [] => rfl - | a :: l => by simp only [flatten, ← map_flatten, ranges_flatten, sum_cons, range_add] + | a :: l => by simp [ranges, ← map_flatten, ranges_flatten, range_add] /-- The members of `l.ranges` have no duplicate -/ theorem ranges_nodup {l s : List ℕ} (hs : s ∈ ranges l) : s.Nodup := diff --git a/Mathlib/Analysis/Convex/StoneSeparation.lean b/Mathlib/Analysis/Convex/StoneSeparation.lean index 5a52d98fda2ea..479beae3a5744 100644 --- a/Mathlib/Analysis/Convex/StoneSeparation.lean +++ b/Mathlib/Analysis/Convex/StoneSeparation.lean @@ -67,8 +67,7 @@ theorem not_disjoint_segment_convexHull_triple {p q u v x y z : E} (hz : z ∈ s ((az * av * bu) • p + ((bz * au * bv) • q + (au * av) • (az • x + bz • y))) · module congr 3 - simp only [smul_add, List.foldr, Fin.reduceFinMk, id_eq, Fin.isValue, Matrix.cons_val_two, - Nat.succ_eq_add_one, Nat.reduceAdd, Matrix.tail_cons, Matrix.head_cons, add_zero, w, z] + simp [w, z] /-- **Stone's Separation Theorem** -/ theorem exists_convex_convex_compl_subset (hs : Convex 𝕜 s) (ht : Convex 𝕜 t) (hst : Disjoint s t) : diff --git a/Mathlib/Computability/TuringMachine.lean b/Mathlib/Computability/TuringMachine.lean index 3af2e1dc86a62..64621943971af 100644 --- a/Mathlib/Computability/TuringMachine.lean +++ b/Mathlib/Computability/TuringMachine.lean @@ -1616,8 +1616,7 @@ theorem stepAux_write (q : Stmt'₁) (v : σ) (a b : Γ) (L R : ListBlank Γ) : cases' l₂' with b l₂' <;> simp only [List.length_nil, List.length_cons, Nat.succ_inj', reduceCtorEq] at e rw [List.reverseAux, ← IH (a :: l₁) l₂' e] - simp only [stepAux, ListBlank.append, Tape.write_mk', Tape.move_right_mk', ListBlank.head_cons, - ListBlank.tail_cons] + simp [stepAux, ListBlank.append, write] variable (encdec : ∀ a, dec (enc a) = a) include encdec diff --git a/Mathlib/GroupTheory/PushoutI.lean b/Mathlib/GroupTheory/PushoutI.lean index c275a0b83ee72..9c027378a289c 100644 --- a/Mathlib/GroupTheory/PushoutI.lean +++ b/Mathlib/GroupTheory/PushoutI.lean @@ -320,9 +320,7 @@ noncomputable def cons {i} (g : G i) (w : NormalWord d) (hmw : w.fstIdx ≠ some @[simp] theorem prod_cons {i} (g : G i) (w : NormalWord d) (hmw : w.fstIdx ≠ some i) (hgr : g ∉ (φ i).range) : (cons g w hmw hgr).prod = of i g * w.prod := by - simp only [prod, cons, Word.prod, List.map, ← of_apply_eq_base φ i, equiv_fst_eq_mul_inv, - mul_assoc, MonoidHom.apply_ofInjective_symm, List.prod_cons, map_mul, map_inv, - ofCoprodI_of, inv_mul_cancel_left] + simp [prod, cons, ← of_apply_eq_base φ i, equiv_fst_eq_mul_inv, mul_assoc] variable [DecidableEq ι] [∀ i, DecidableEq (G i)] diff --git a/Mathlib/ModelTheory/Algebra/Ring/Definability.lean b/Mathlib/ModelTheory/Algebra/Ring/Definability.lean index 216794ac84fa0..2d3feb0fc21bc 100644 --- a/Mathlib/ModelTheory/Algebra/Ring/Definability.lean +++ b/Mathlib/ModelTheory/Algebra/Ring/Definability.lean @@ -38,11 +38,7 @@ theorem mvPolynomial_zeroLocus_definable {ι K : Type*} [Field K] (Sum.map (fun p => ⟨p.1.1.coeff p.2.1, by simp only [Set.mem_iUnion] exact ⟨p.1.1, p.1.2, Set.mem_image_of_mem _ p.2.2⟩⟩) id)) 0), ?_⟩ - simp only [Finset.mem_coe, Formula.Realize, Term.equal, Term.relabel_relabel, Function.comp_def, - Term.relabel, realize_iInf, Finset.mem_attach, realize_bdEqual, Term.realize_relabel, - Sum.elim_inl, realize_termOfFreeCommRing, lift_genericPolyMap, Sum.map_inr, id_eq, Sum.elim_inr, - Sum.map_inl, MvPolynomialSupportLEEquiv_symm_apply_coeff, Term.realize_func, Sum.elim_comp_inl, - CompatibleRing.funMap_zero, true_implies, Subtype.forall, p'] + simp [Formula.Realize, Term.equal, Function.comp_def, p'] end Ring diff --git a/Mathlib/ModelTheory/Syntax.lean b/Mathlib/ModelTheory/Syntax.lean index ff03d55615d75..7461119932349 100644 --- a/Mathlib/ModelTheory/Syntax.lean +++ b/Mathlib/ModelTheory/Syntax.lean @@ -642,7 +642,7 @@ theorem comp_onBoundedFormula {L'' : Language} (φ : L' →ᴸ L'') (ψ : L → ext f induction f with | falsum => rfl - | equal => simp only [onBoundedFormula, comp_onTerm, Function.comp_apply] + | equal => simp [Term.bdEqual] | rel => simp only [onBoundedFormula, comp_onRelation, comp_onTerm, Function.comp_apply]; rfl | imp _ _ ih1 ih2 => simp only [onBoundedFormula, Function.comp_apply, ih1, ih2, eq_self_iff_true, and_self_iff] diff --git a/Mathlib/Testing/Plausible/Functions.lean b/Mathlib/Testing/Plausible/Functions.lean index ba4209766edf6..a8801b6e6b2ad 100644 --- a/Mathlib/Testing/Plausible/Functions.lean +++ b/Mathlib/Testing/Plausible/Functions.lean @@ -194,8 +194,7 @@ theorem List.applyId_zip_eq [DecidableEq α] {xs ys : List α} (h₀ : List.Nodu subst h₂ cases ys · cases h₁ - · simp only [applyId, map, Prod.toSigma, dlookup_cons_eq, Option.getD_some, - getElem?_cons_zero, Option.some.injEq] + · simp · cases ys · cases h₁ · cases' h₀ with _ _ h₀ h₁ From 01e0745f82a82d9c74042320a342038cd49ba499 Mon Sep 17 00:00:00 2001 From: Mitchell Lee Date: Fri, 20 Dec 2024 04:08:44 +0000 Subject: [PATCH 797/829] feat: prime divisors of a prime power (#20095) If a prime $p$ divides a prime power $q^m$, then $p = q$. The proof is a shortening of the one that Ruben Van de Velde posted here: https://leanprover.zulipchat.com/#narrow/channel/217875-Is-there-code-for-X.3F/topic/.E2.9C.94.20If.20prime.20divides.20power.20of.20prime.2C.20then.20both.20primes.20equal Co-authored-by: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> --- Mathlib/Data/Nat/Prime/Basic.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/Data/Nat/Prime/Basic.lean b/Mathlib/Data/Nat/Prime/Basic.lean index e1452b5e145a5..8cb621b8201d1 100644 --- a/Mathlib/Data/Nat/Prime/Basic.lean +++ b/Mathlib/Data/Nat/Prime/Basic.lean @@ -194,6 +194,9 @@ theorem eq_or_coprime_of_le_prime {n p} (n_pos : 0 < n) (hle : n ≤ p) (pp : Pr p = n ∨ Coprime p n := hle.eq_or_lt.imp Eq.symm fun h => coprime_of_lt_prime n_pos h pp +theorem prime_eq_prime_of_dvd_pow {m p q} (pp : Prime p) (pq : Prime q) (h : p ∣ q ^ m) : p = q := + (prime_dvd_prime_iff_eq pp pq).mp (pp.dvd_of_dvd_pow h) + theorem dvd_prime_pow {p : ℕ} (pp : Prime p) {m i : ℕ} : i ∣ p ^ m ↔ ∃ k ≤ m, i = p ^ k := by simp_rw [_root_.dvd_prime_pow (prime_iff.mp pp) m, associated_eq_eq] From 4e34721f6c89c3ff4282b58103694bbb523c209f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Fri, 20 Dec 2024 04:35:20 +0000 Subject: [PATCH 798/829] feat(Algebra/Homology/Embedding): the homology of an extension of homological complexes (#19198) Given an embedding `e : c.Embedding c'`, `K : HomologicalComplex C c`, and degrees such that `e.f j = j'`, we compute the homology of `K.extend e` in degree `j'` in terms of the homology of `K` in degree `j`. --- .../Homology/Embedding/ExtendHomology.lean | 135 ++++++++++++++++++ 1 file changed, 135 insertions(+) diff --git a/Mathlib/Algebra/Homology/Embedding/ExtendHomology.lean b/Mathlib/Algebra/Homology/Embedding/ExtendHomology.lean index 123c91edcc3e7..200b3f46bb39a 100644 --- a/Mathlib/Algebra/Homology/Embedding/ExtendHomology.lean +++ b/Mathlib/Algebra/Homology/Embedding/ExtendHomology.lean @@ -44,6 +44,18 @@ lemma comp_d_eq_zero_iff ⦃W : C⦄ (φ : W ⟶ K.X j) : rw [hk] exact hjk +include hi hi' in +lemma d_comp_eq_zero_iff ⦃W : C⦄ (φ : K.X j ⟶ W) : + K.d i j ≫ φ = 0 ↔ (K.extend e).d i' j' ≫ (K.extendXIso e hj').hom ≫ φ = 0 := by + by_cases hij : c.Rel i j + · have hi' : e.f i = i' := by rw [← hi', ← hj', c'.prev_eq' (e.rel hij)] + rw [K.extend_d_eq e hi' hj', assoc, assoc, Iso.inv_hom_id_assoc, + ← cancel_epi (K.extendXIso e hi').hom, comp_zero] + · simp only [K.shape _ _ hij, zero_comp, true_iff] + rw [K.extend_d_to_eq_zero e i' j' j hj', zero_comp] + rw [hi] + exact hij + namespace leftHomologyData variable (cone : KernelFork (K.d j k)) (hcone : IsLimit cone) @@ -136,8 +148,131 @@ noncomputable def leftHomologyData (h : (K.sc' i j k).LeftHomologyData) : exact h.wπ hπ := isColimitCokernelCofork K e hj' hi hi' hk hk' _ h.hi _ h.hπ +namespace rightHomologyData + +variable (cocone : CokernelCofork (K.d i j)) (hcocone : IsColimit cocone) + +/-- The cokernel cofork of `(K.extend e).d i' j'` that is deduced from a cokernel +cofork of `K.d i j`. -/ +@[simp] +noncomputable def cokernelCofork : CokernelCofork ((K.extend e).d i' j') := + CokernelCofork.ofπ ((extendXIso K e hj').hom ≫ cocone.π) (by + rw [← d_comp_eq_zero_iff K e hj' hi hi' cocone.π, cocone.condition]) + +/-- The colimit cokernel cofork of `(K.extend e).d i' j'` that is deduced from a +colimit cokernel cofork of `K.d i j`. -/ +noncomputable def isColimitCokernelCofork : IsColimit (cokernelCofork K e hj' hi hi' cocone) := + CokernelCofork.isColimitOfIsColimitOfIff hcocone ((K.extend e).d i' j') + (extendXIso K e hj') (d_comp_eq_zero_iff K e hj' hi hi') + +variable (cone : KernelFork (hcocone.desc (CokernelCofork.ofπ (K.d j k) (K.d_comp_d i j k)))) + (hcone : IsLimit cone) + +include hk hk' hcocone in +lemma d_comp_desc_eq_zero_iff' ⦃W : C⦄ (f' : cocone.pt ⟶ K.X k) + (hf' : cocone.π ≫ f' = K.d j k) + (f'' : cocone.pt ⟶ (K.extend e).X k') + (hf'' : (extendXIso K e hj').hom ≫ cocone.π ≫ f'' = (K.extend e).d j' k') + (φ : W ⟶ cocone.pt) : + φ ≫ f' = 0 ↔ φ ≫ f'' = 0 := by + by_cases hjk : c.Rel j k + · have hk'' : e.f k = k' := by rw [← hk', ← hj', c'.next_eq' (e.rel hjk)] + have : f' ≫ (K.extendXIso e hk'').inv = f'' := by + apply Cofork.IsColimit.hom_ext hcocone + rw [reassoc_of% hf', ← cancel_epi (extendXIso K e hj').hom, hf'', + K.extend_d_eq e hj' hk''] + rw [← cancel_mono (K.extendXIso e hk'').inv, zero_comp, assoc, this] + · have h₁ : f' = 0 := by + apply Cofork.IsColimit.hom_ext hcocone + simp only [hf', comp_zero, K.shape _ _ hjk] + have h₂ : f'' = 0 := by + apply Cofork.IsColimit.hom_ext hcocone + rw [← cancel_epi (extendXIso K e hj').hom, hf'', comp_zero, comp_zero, + K.extend_d_from_eq_zero e j' k' j hj'] + rw [hk] + exact hjk + simp [h₁, h₂] + +include hk hk' in +lemma d_comp_desc_eq_zero_iff ⦃W : C⦄ (φ : W ⟶ cocone.pt) : + φ ≫ hcocone.desc (CokernelCofork.ofπ (K.d j k) (K.d_comp_d i j k)) = 0 ↔ + φ ≫ ((isColimitCokernelCofork K e hj' hi hi' cocone hcocone).desc + (CokernelCofork.ofπ ((K.extend e).d j' k') (d_comp_d _ _ _ _))) = 0 := + d_comp_desc_eq_zero_iff' K e hj' hk hk' cocone hcocone _ (hcocone.fac _ _) _ (by + simpa using (isColimitCokernelCofork K e hj' hi hi' cocone hcocone).fac _ + WalkingParallelPair.one) _ + +/-- Auxiliary definition for `extend.rightHomologyData`. -/ +noncomputable def kernelFork : + KernelFork ((isColimitCokernelCofork K e hj' hi hi' cocone hcocone).desc + (CokernelCofork.ofπ ((K.extend e).d j' k') (d_comp_d _ _ _ _))) := + KernelFork.ofι cone.ι (by + rw [← d_comp_desc_eq_zero_iff K e hj' hi hi' hk hk' cocone hcocone] + exact cone.condition) + +/-- Auxiliary definition for `extend.rightHomologyData`. -/ +noncomputable def isLimitKernelFork : + IsLimit (kernelFork K e hj' hi hi' hk hk' cocone hcocone cone) := + KernelFork.isLimitOfIsLimitOfIff' hcone _ + (d_comp_desc_eq_zero_iff K e hj' hi hi' hk hk' cocone hcocone) + +end rightHomologyData + +open rightHomologyData in +/-- The right homology data of `(K.extend e).sc' i' j' k'` that is deduced +from a right homology data of `K.sc' i j k`. -/ +@[simps] +noncomputable def rightHomologyData (h : (K.sc' i j k).RightHomologyData) : + ((K.extend e).sc' i' j' k').RightHomologyData where + Q := h.Q + H := h.H + p := (extendXIso K e hj').hom ≫ h.p + ι := h.ι + wp := by + dsimp + rw [← d_comp_eq_zero_iff K e hj' hi hi'] + exact h.wp + hp := isColimitCokernelCofork K e hj' hi hi' _ h.hp + wι := by + dsimp + rw [← d_comp_desc_eq_zero_iff K e hj' hi hi' hk hk' _ h.hp] + exact h.wι + hι := isLimitKernelFork K e hj' hi hi' hk hk' _ h.hp _ h.hι + +/-- The homology data of `(K.extend e).sc' i' j' k'` that is deduced +from a homology data of `K.sc' i j k`. -/ +@[simps] +noncomputable def homologyData (h : (K.sc' i j k).HomologyData) : + ((K.extend e).sc' i' j' k').HomologyData where + left := leftHomologyData K e hj' hi hi' hk hk' h.left + right := rightHomologyData K e hj' hi hi' hk hk' h.right + iso := h.iso + +/-- The homology data of `(K.extend e).sc j'` that is deduced +from a homology data of `K.sc' i j k`. -/ +@[simps!] +noncomputable def homologyData' (h : (K.sc' i j k).HomologyData) : + ((K.extend e).sc j').HomologyData := + homologyData K e hj' hi rfl hk rfl h + end HomologyData +lemma hasHomology {j : ι} {j' : ι'} (hj' : e.f j = j') [K.HasHomology j] : + (K.extend e).HasHomology j' := + ShortComplex.HasHomology.mk' + (homologyData' K e hj' rfl rfl ((K.sc j).homologyData)) + +instance (j : ι) [K.HasHomology j] : (K.extend e).HasHomology (e.f j) := + hasHomology K e rfl + +instance [∀ j, K.HasHomology j] (j' : ι') : (K.extend e).HasHomology j' := by + by_cases h : ∃ j, e.f j = j' + · obtain ⟨j, rfl⟩ := h + infer_instance + · have hj := isZero_extend_X K e j' (by tauto) + exact ShortComplex.HasHomology.mk' + (ShortComplex.HomologyData.ofZeros _ (hj.eq_of_tgt _ _) (hj.eq_of_src _ _)) + end extend end HomologicalComplex From 86d6ee4be09b8be1c2fde741ddac318fdccf49b5 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Fri, 20 Dec 2024 04:35:21 +0000 Subject: [PATCH 799/829] refactor(LinearAlgebra/Matrix/FixedDetMatrices): Inline `S_T_subgroup` (#20086) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Per @jcommelin's comment on #15851, I think it's better to inline `S_T_subgroup` since it has a short definition (`closure {S, T}`) ends up being equal to `⊤`. --- .../Matrix/FixedDetMatrices.lean | 38 +++++++------------ 1 file changed, 13 insertions(+), 25 deletions(-) diff --git a/Mathlib/LinearAlgebra/Matrix/FixedDetMatrices.lean b/Mathlib/LinearAlgebra/Matrix/FixedDetMatrices.lean index 7c9dbfd95a909..834276a8fdccd 100644 --- a/Mathlib/LinearAlgebra/Matrix/FixedDetMatrices.lean +++ b/Mathlib/LinearAlgebra/Matrix/FixedDetMatrices.lean @@ -244,35 +244,12 @@ theorem induction_on {C : Δ m → Prop} {A : Δ m} (hm : m ≠ 0) rw [← reduce_reduceStep hc] at hA simpa only [reduceStep, prop_red_S hS, prop_red_T_pow hS hT] using ih hA -/--The subgroup of `SL(2,ℤ)` generated by `S` and `T`. -/ -def S_T_subgroup := Subgroup.closure {S, T} - -lemma S_mem_S_T_subgroup : S ∈ S_T_subgroup := by - apply Subgroup.subset_closure - simp only [Set.mem_insert_iff, Set.mem_singleton_iff, true_or] - -lemma T_mem_S_T_subgroup : T ∈ S_T_subgroup := by - apply Subgroup.subset_closure - simp only [Set.mem_insert_iff, Set.mem_singleton_iff, or_true] - lemma reps_one_id (A : FixedDetMatrix (Fin 2) ℤ 1) (a1 : A.1 1 0 = 0) (a4 : 0 < A.1 0 0) (a6 : |A.1 0 1| < |(A.1 1 1)|) : A = (1 : SL(2, ℤ)) := by have := Int.mul_eq_one_iff_eq_one_or_neg_one.mp (A_c_eq_zero a1) ext i j fin_cases i <;> fin_cases j <;> aesop -lemma det_one_mem_S_T_subgroup (A : Δ 1) : A ∈ S_T_subgroup := by - induction A using (induction_on Int.one_ne_zero) with - | h0 A a1 a4 _ a6 => - rw [reps_one_id A a1 a4 a6] - exact Subgroup.one_mem S_T_subgroup - | hS B hb => - rw [smul_eq_mul] - apply Subgroup.mul_mem _ (S_mem_S_T_subgroup) (hb) - | hT B hb => - rw [smul_eq_mul] - apply Subgroup.mul_mem _ (T_mem_S_T_subgroup) (hb) - end IntegralFixedDetMatrices end FixedDetMatrices @@ -281,8 +258,19 @@ open MatrixGroups FixedDetMatrices section SL2Z_generators +open ModularGroup Subgroup + /-- `SL(2, ℤ)` is generated by `S` and `T`. -/ -lemma SpecialLinearGroup.SL2Z_generators : S_T_subgroup = ⊤ := by - refine (Subgroup.eq_top_iff' S_T_subgroup).mpr (fun _ => det_one_mem_S_T_subgroup _) +lemma SpecialLinearGroup.SL2Z_generators : closure {S, T} = ⊤ := by + rw [eq_top_iff'] + intro A + induction A using (induction_on one_ne_zero) with + | h0 A a1 a4 _ a6 => + rw [reps_one_id A a1 a4 a6] + exact one_mem _ + | hS B hb => + exact mul_mem (subset_closure (Set.mem_insert S {T})) hb + | hT B hb => + exact mul_mem (subset_closure (Set.mem_insert_of_mem S rfl)) hb end SL2Z_generators From 84c51e92ee7de31a3f942e076cf5855266823d6b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Martin=20Dvo=C5=99=C3=A1k?= Date: Fri, 20 Dec 2024 04:44:21 +0000 Subject: [PATCH 800/829] docs(LinearAlgebra/Matrix/Determinant/Basic): det_of_*erTriangular (#20072) --- Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean index 544a0f20ffac2..230982afbb29a 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean @@ -633,7 +633,7 @@ theorem det_blockDiagonal {o : Type*} [Fintype o] [DecidableEq o] (M : o → Mat /-- The determinant of a 2×2 block matrix with the lower-left block equal to zero is the product of the determinants of the diagonal blocks. For the generalization to any number of blocks, see -`Matrix.det_of_upper_triangular`. -/ +`Matrix.det_of_upperTriangular`. -/ @[simp] theorem det_fromBlocks_zero₂₁ (A : Matrix m m R) (B : Matrix m n R) (D : Matrix n n R) : (Matrix.fromBlocks A B 0 D).det = A.det * D.det := by @@ -687,7 +687,7 @@ theorem det_fromBlocks_zero₂₁ (A : Matrix m m R) (B : Matrix m n R) (D : Mat /-- The determinant of a 2×2 block matrix with the upper-right block equal to zero is the product of the determinants of the diagonal blocks. For the generalization to any number of blocks, see -`Matrix.det_of_lower_triangular`. -/ +`Matrix.det_of_lowerTriangular`. -/ @[simp] theorem det_fromBlocks_zero₁₂ (A : Matrix m m R) (C : Matrix n m R) (D : Matrix n n R) : (Matrix.fromBlocks A 0 C D).det = A.det * D.det := by From 75a1c934df546f46c1dbf72eda4db069da6ba10b Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 20 Dec 2024 04:44:22 +0000 Subject: [PATCH 801/829] chore(Coxeter/Inversion): use `getElem?` (#20104) --- Mathlib/GroupTheory/Coxeter/Inversion.lean | 19 +++++++++---------- 1 file changed, 9 insertions(+), 10 deletions(-) diff --git a/Mathlib/GroupTheory/Coxeter/Inversion.lean b/Mathlib/GroupTheory/Coxeter/Inversion.lean index 43a790543f337..4603c629ae971 100644 --- a/Mathlib/GroupTheory/Coxeter/Inversion.lean +++ b/Mathlib/GroupTheory/Coxeter/Inversion.lean @@ -245,7 +245,7 @@ theorem leftInvSeq_reverse (ω : List B) : theorem getD_rightInvSeq (ω : List B) (j : ℕ) : (ris ω).getD j 1 = (π (ω.drop (j + 1)))⁻¹ - * (Option.map (cs.simple) (ω.get? j)).getD 1 + * (Option.map (cs.simple) ω[j]?).getD 1 * π (ω.drop (j + 1)) := by induction' ω with i ω ih generalizing j · simp @@ -258,14 +258,14 @@ theorem getD_rightInvSeq (ω : List B) (j : ℕ) : lemma getElem_rightInvSeq (ω : List B) (j : ℕ) (h : j < ω.length) : (ris ω)[j]'(by simp[h]) = (π (ω.drop (j + 1)))⁻¹ - * (Option.map (cs.simple) (ω.get? j)).getD 1 + * (Option.map (cs.simple) ω[j]?).getD 1 * π (ω.drop (j + 1)) := by rw [← List.getD_eq_getElem (ris ω) 1, getD_rightInvSeq] theorem getD_leftInvSeq (ω : List B) (j : ℕ) : (lis ω).getD j 1 = π (ω.take j) - * (Option.map (cs.simple) (ω.get? j)).getD 1 + * (Option.map (cs.simple) ω[j]?).getD 1 * (π (ω.take j))⁻¹ := by induction' ω with i ω ih generalizing j · simp @@ -288,18 +288,18 @@ theorem getD_rightInvSeq_mul_self (ω : List B) (j : ℕ) : ((ris ω).getD j 1) * ((ris ω).getD j 1) = 1 := by simp_rw [getD_rightInvSeq, mul_assoc] rcases em (j < ω.length) with hj | nhj - · rw [get?_eq_get hj] + · rw [getElem?_eq_getElem hj] simp [← mul_assoc] - · rw [get?_eq_none_iff.mpr (by omega)] + · rw [getElem?_eq_none_iff.mpr (by omega)] simp theorem getD_leftInvSeq_mul_self (ω : List B) (j : ℕ) : ((lis ω).getD j 1) * ((lis ω).getD j 1) = 1 := by simp_rw [getD_leftInvSeq, mul_assoc] rcases em (j < ω.length) with hj | nhj - · rw [get?_eq_get hj] + · rw [getElem?_eq_getElem hj] simp [← mul_assoc] - · rw [get?_eq_none_iff.mpr (by omega)] + · rw [getElem?_eq_none_iff.mpr (by omega)] simp theorem rightInvSeq_drop (ω : List B) (j : ℕ) : @@ -417,9 +417,8 @@ theorem IsReduced.nodup_rightInvSeq {ω : List B} (rω : cs.IsReduced ω) : List drop_drop, nil_append, min_eq_left_of_lt (j_lt_j'.trans j'_lt_length), Nat.add_comm, ← add_assoc, Nat.sub_add_cancel (by omega), mul_left_inj, mul_right_inj] congr 2 - show get? (take j ω ++ drop (j + 1) ω) (j' - 1) = get? ω j' - rw [get?_eq_getElem?, get?_eq_getElem?, - getElem?_append_right (by simp [Nat.le_sub_one_of_lt j_lt_j']), getElem?_drop] + show (take j ω ++ drop (j + 1) ω)[j' - 1]? = ω[j']? + rw [getElem?_append_right (by simp [Nat.le_sub_one_of_lt j_lt_j']), getElem?_drop] congr show j + 1 + (j' - 1 - List.length (take j ω)) = j' rw [length_take] From 7a9aa7e228327a9ee3c96c29961ab6e0b0ce0cad Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Fri, 20 Dec 2024 05:03:47 +0000 Subject: [PATCH 802/829] feat(Algebra/Group/Subgroup/Ker): Add `map_lt_map_iff_of_injective` (#19826) This PR adds `lt` versions of `map_le_map_iff_of_injective` and `map_subtype_le_map_subtype`. Co-authored-by: Thomas Browning --- Mathlib/Algebra/Group/Subgroup/Ker.lean | 12 +++++++++++- 1 file changed, 11 insertions(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Group/Subgroup/Ker.lean b/Mathlib/Algebra/Group/Subgroup/Ker.lean index 8fd902758f7c1..c0c7e47007b41 100644 --- a/Mathlib/Algebra/Group/Subgroup/Ker.lean +++ b/Mathlib/Algebra/Group/Subgroup/Ker.lean @@ -460,7 +460,17 @@ theorem map_le_map_iff_of_injective {f : G →* N} (hf : Function.Injective f) { @[to_additive (attr := simp)] theorem map_subtype_le_map_subtype {G' : Subgroup G} {H K : Subgroup G'} : H.map G'.subtype ≤ K.map G'.subtype ↔ H ≤ K := - map_le_map_iff_of_injective <| by apply Subtype.coe_injective + map_le_map_iff_of_injective G'.subtype_injective + +@[to_additive] +theorem map_lt_map_iff_of_injective {f : G →* N} (hf : Function.Injective f) {H K : Subgroup G} : + H.map f < K.map f ↔ H < K := + lt_iff_lt_of_le_iff_le' (map_le_map_iff_of_injective hf) (map_le_map_iff_of_injective hf) + +@[to_additive (attr := simp)] +theorem map_subtype_lt_map_subtype {G' : Subgroup G} {H K : Subgroup G'} : + H.map G'.subtype < K.map G'.subtype ↔ H < K := + map_lt_map_iff_of_injective G'.subtype_injective @[to_additive] theorem map_injective {f : G →* N} (h : Function.Injective f) : Function.Injective (map f) := From 899ea23aae63245bb91703cb6affde7f62276400 Mon Sep 17 00:00:00 2001 From: David Renshaw Date: Fri, 20 Dec 2024 05:03:48 +0000 Subject: [PATCH 803/829] chore(Data/List/Lemmas): deprecate tail_reverse_eq_reverse_dropLast (#19874) Makes `tail_reverse_eq_reverse_dropLast` a deprecated alias to [`tail_reverse`](https://github.com/leanprover/lean4/blob/938651121f3d8819109eee75722a1de087feff71/src/Init/Data/List/Lemmas.lean#L2990-L2994), which was added to core in https://github.com/leanprover/lean4/pull/5360. This duplicated lemma was found by [`tryAtEachStep`](https://github.com/dwrensha/tryAtEachStep). --- Mathlib/Data/List/Lemmas.lean | 11 +---------- 1 file changed, 1 insertion(+), 10 deletions(-) diff --git a/Mathlib/Data/List/Lemmas.lean b/Mathlib/Data/List/Lemmas.lean index a6c98a38f2db9..ef62ad8d898a9 100644 --- a/Mathlib/Data/List/Lemmas.lean +++ b/Mathlib/Data/List/Lemmas.lean @@ -31,16 +31,7 @@ lemma count_flatMap [BEq β] (l : List α) (f : α → List β) (x : β) : @[deprecated (since := "2024-08-20")] alias getElem_reverse' := getElem_reverse -theorem tail_reverse_eq_reverse_dropLast (l : List α) : - l.reverse.tail = l.dropLast.reverse := by - ext i v; by_cases hi : i < l.length - 1 - · simp only [← drop_one] - rw [getElem?_eq_getElem (by simpa), getElem?_eq_getElem (by simpa), - ← getElem_drop' _, getElem_reverse, getElem_reverse, getElem_dropLast] - · simp [show l.length - 1 - (1 + i) = l.length - 1 - 1 - i by omega] - all_goals ((try simp); omega) - · rw [getElem?_eq_none, getElem?_eq_none] - all_goals (simp; omega) +@[deprecated (since := "2024-12-10")] alias tail_reverse_eq_reverse_dropLast := tail_reverse @[deprecated (since := "2024-08-19")] alias nthLe_tail := getElem_tail From cfc731aeac09a4a42eafd783067c5c3164f7b124 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Fri, 20 Dec 2024 05:03:49 +0000 Subject: [PATCH 804/829] feat(GroupTheory/SpecificGroups/ZGroup): a finite Z-group has cyclic abelianization (#19880) This is one step towards showing that a finite Z-group `G` is the semidirect product of the commutator subgroup `G'` and the abelianziation `G/G'` and that `G'` and `G/G'` are cyclic of coprime orders. --- .../GroupTheory/SpecificGroups/Cyclic.lean | 3 +++ .../GroupTheory/SpecificGroups/ZGroup.lean | 20 +++++++++++++++---- 2 files changed, 19 insertions(+), 4 deletions(-) diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index ba18539a63662..04d3cf186891d 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -81,6 +81,9 @@ def IsCyclic.commGroup [hg : Group α] [IsCyclic α] : CommGroup α := let ⟨_, hm⟩ := hg y hm ▸ hn ▸ zpow_mul_comm _ _ _ } +instance [Group G] (H : Subgroup G) [IsCyclic H] : H.IsCommutative := + ⟨⟨IsCyclic.commGroup.mul_comm⟩⟩ + variable [Group α] [Group G] [Group G'] /-- A non-cyclic multiplicative group is non-trivial. -/ diff --git a/Mathlib/GroupTheory/SpecificGroups/ZGroup.lean b/Mathlib/GroupTheory/SpecificGroups/ZGroup.lean index f2d850b5a8d40..d2b1094c66daa 100644 --- a/Mathlib/GroupTheory/SpecificGroups/ZGroup.lean +++ b/Mathlib/GroupTheory/SpecificGroups/ZGroup.lean @@ -13,7 +13,11 @@ A Z-group is a group whose Sylow subgroups are all cyclic. ## Main definitions -* `IsZGroup G`: A predicate stating that all Sylow subgroups of `G` are cyclic. +* `IsZGroup G`: a predicate stating that all Sylow subgroups of `G` are cyclic. + +## Main results + +* `IsZGroup.isCyclic_abelianization`: a finite Z-group has cyclic abelianization. TODO: Show that if `G` is a Z-group with commutator subgroup `G'`, then `G = G' ⋊ G/G'` where `G'` and `G/G'` are cyclic of coprime orders. @@ -62,6 +66,9 @@ theorem of_surjective [Finite G] [hG : IsZGroup G] (hf : Function.Surjective f) instance [Finite G] [IsZGroup G] (H : Subgroup G) [H.Normal] : IsZGroup (G ⧸ H) := of_surjective (QuotientGroup.mk'_surjective H) +section Nilpotent + +variable (G) in theorem exponent_eq_card [Finite G] [IsZGroup G] : Monoid.exponent G = Nat.card G := by refine dvd_antisymm Group.exponent_dvd_nat_card ?_ rw [← Nat.factorization_prime_le_iff_dvd Nat.card_pos.ne' Monoid.exponent_ne_zero_of_finite] @@ -75,11 +82,16 @@ theorem exponent_eq_card [Finite G] [IsZGroup G] : Monoid.exponent G = Nat.card instance [Finite G] [IsZGroup G] [hG : Group.IsNilpotent G] : IsCyclic G := by have (p : { x // x ∈ (Nat.card G).primeFactors }) : Fact p.1.Prime := ⟨Nat.prime_of_mem_primeFactors p.2⟩ - let h (p : { x // x ∈ (Nat.card G).primeFactors }) (P : Sylow p G) : CommGroup P := - IsCyclic.commGroup obtain ⟨ϕ⟩ := ((isNilpotent_of_finite_tfae (G := G)).out 0 4).mp hG let _ : CommGroup G := ⟨fun g h ↦ by rw [← ϕ.symm.injective.eq_iff, map_mul, mul_comm, ← map_mul]⟩ - exact IsCyclic.of_exponent_eq_card exponent_eq_card + exact IsCyclic.of_exponent_eq_card (exponent_eq_card G) + +/-- A finite Z-group has cyclic abelianization. -/ +instance isCyclic_abelianization [Finite G] [IsZGroup G] : IsCyclic (Abelianization G) := + let _ : IsZGroup (Abelianization G) := inferInstanceAs (IsZGroup (G ⧸ commutator G)) + inferInstance + +end Nilpotent end IsZGroup From d23ee2725c60f20116fbf43661187f9b25cb74c5 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Fri, 20 Dec 2024 05:03:50 +0000 Subject: [PATCH 805/829] feat(GroupTheory/Abelianization): Add `Abelianization.ker_of` (#19906) This PR adds `Abelianization.ker_of`. --- Mathlib/GroupTheory/Abelianization.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Mathlib/GroupTheory/Abelianization.lean b/Mathlib/GroupTheory/Abelianization.lean index 0d35f5937e6b3..e27a628bddbad 100644 --- a/Mathlib/GroupTheory/Abelianization.lean +++ b/Mathlib/GroupTheory/Abelianization.lean @@ -112,6 +112,11 @@ def of : G →* Abelianization G where theorem mk_eq_of (a : G) : Quot.mk _ a = of a := rfl +variable (G) in +@[simp] +theorem ker_of : of.ker = commutator G := + QuotientGroup.ker_mk' (commutator G) + section lift -- So far we have built Gᵃᵇ and proved it's an abelian group. From 024a16f8c6904b7761010cf09ddaf17200646b87 Mon Sep 17 00:00:00 2001 From: blizzard_inc Date: Fri, 20 Dec 2024 05:03:51 +0000 Subject: [PATCH 806/829] chore(Data/List/AList): Rename `insert_of_neg` and similar to `insert_of_not_mem` (#19960) completes a todo, also mentioned in #7987 also deprecates the old things --- Mathlib/Data/Finmap.lean | 8 +++++--- Mathlib/Data/List/AList.lean | 26 +++++++++++++++----------- 2 files changed, 20 insertions(+), 14 deletions(-) diff --git a/Mathlib/Data/Finmap.lean b/Mathlib/Data/Finmap.lean index 2b02b71dd05c5..93f83e0f469b8 100644 --- a/Mathlib/Data/Finmap.lean +++ b/Mathlib/Data/Finmap.lean @@ -417,11 +417,13 @@ theorem insert_toFinmap (a : α) (b : β a) (s : AList β) : insert a b (AList.toFinmap s) = AList.toFinmap (s.insert a b) := by simp [insert] -theorem insert_entries_of_neg {a : α} {b : β a} {s : Finmap β} : +theorem entries_insert_of_not_mem {a : α} {b : β a} {s : Finmap β} : a ∉ s → (insert a b s).entries = ⟨a, b⟩ ::ₘ s.entries := induction_on s fun s h => by - -- Porting note: `-insert_entries` required - simp [AList.insert_entries_of_neg (mt mem_toFinmap.1 h), -insert_entries] + -- Porting note: `-entries_insert` required + simp [AList.entries_insert_of_not_mem (mt mem_toFinmap.1 h), -entries_insert] + +@[deprecated (since := "2024-12-14")] alias insert_entries_of_neg := entries_insert_of_not_mem @[simp] theorem mem_insert {a a' : α} {b' : β a'} {s : Finmap β} : a ∈ insert a' b' s ↔ a = a' ∨ a ∈ s := diff --git a/Mathlib/Data/List/AList.lean b/Mathlib/Data/List/AList.lean index cb4442a55a060..c9227ba6d8e96 100644 --- a/Mathlib/Data/List/AList.lean +++ b/Mathlib/Data/List/AList.lean @@ -237,17 +237,21 @@ def insert (a : α) (b : β a) (s : AList β) : AList β := ⟨kinsert a b s.entries, kinsert_nodupKeys a b s.nodupKeys⟩ @[simp] -theorem insert_entries {a} {b : β a} {s : AList β} : +theorem entries_insert {a} {b : β a} {s : AList β} : (insert a b s).entries = Sigma.mk a b :: kerase a s.entries := rfl -theorem insert_entries_of_neg {a} {b : β a} {s : AList β} (h : a ∉ s) : - (insert a b s).entries = ⟨a, b⟩ :: s.entries := by rw [insert_entries, kerase_of_not_mem_keys h] +@[deprecated (since := "2024-12-17")] alias insert_entries := entries_insert --- Todo: rename to `insert_of_not_mem`. -theorem insert_of_neg {a} {b : β a} {s : AList β} (h : a ∉ s) : +theorem entries_insert_of_not_mem {a} {b : β a} {s : AList β} (h : a ∉ s) : + (insert a b s).entries = ⟨a, b⟩ :: s.entries := by rw [entries_insert, kerase_of_not_mem_keys h] + +theorem insert_of_not_mem {a} {b : β a} {s : AList β} (h : a ∉ s) : insert a b s = ⟨⟨a, b⟩ :: s.entries, nodupKeys_cons.2 ⟨h, s.2⟩⟩ := - ext <| insert_entries_of_neg h + ext <| entries_insert_of_not_mem h + +@[deprecated (since := "2024-12-14")] alias insert_entries_of_neg := entries_insert_of_not_mem +@[deprecated (since := "2024-12-14")] alias insert_of_neg := insert_of_not_mem @[simp] theorem insert_empty (a) (b : β a) : insert a b ∅ = singleton a b := @@ -263,7 +267,7 @@ theorem keys_insert {a} {b : β a} (s : AList β) : (insert a b s).keys = a :: s theorem perm_insert {a} {b : β a} {s₁ s₂ : AList β} (p : s₁.entries ~ s₂.entries) : (insert a b s₁).entries ~ (insert a b s₂).entries := by - simp only [insert_entries]; exact p.kinsert s₁.nodupKeys + simp only [entries_insert]; exact p.kinsert s₁.nodupKeys @[simp] theorem lookup_insert {a} {b : β a} (s : AList β) : lookup a (insert a b s) = some b := by @@ -287,17 +291,17 @@ theorem lookup_to_alist {a} (s : List (Sigma β)) : lookup a s.toAList = s.dlook @[simp] theorem insert_insert {a} {b b' : β a} (s : AList β) : (s.insert a b).insert a b' = s.insert a b' := by - ext : 1; simp only [AList.insert_entries, List.kerase_cons_eq] + ext : 1; simp only [AList.entries_insert, List.kerase_cons_eq] theorem insert_insert_of_ne {a a'} {b : β a} {b' : β a'} (s : AList β) (h : a ≠ a') : ((s.insert a b).insert a' b').entries ~ ((s.insert a' b').insert a b).entries := by - simp only [insert_entries]; rw [kerase_cons_ne, kerase_cons_ne, kerase_comm] <;> + simp only [entries_insert]; rw [kerase_cons_ne, kerase_cons_ne, kerase_comm] <;> [apply Perm.swap; exact h; exact h.symm] @[simp] theorem insert_singleton_eq {a : α} {b b' : β a} : insert a b (singleton a b') = singleton a b := ext <| by - simp only [AList.insert_entries, List.kerase_cons_eq, and_self_iff, AList.singleton_entries, + simp only [AList.entries_insert, List.kerase_cons_eq, and_self_iff, AList.singleton_entries, heq_iff_eq, eq_self_iff_true] @[simp] @@ -342,7 +346,7 @@ theorem insertRec_insert {C : AList β → Sort*} (H0 : C ∅) (IH c.1 c.2 ⟨l, hl⟩ h (@insertRec α β _ C H0 IH ⟨l, hl⟩)) by cases c apply eq_of_heq - convert this <;> rw [insert_of_neg h] + convert this <;> rw [insert_of_not_mem h] rw [insertRec] apply cast_heq From e98084886af394454d85a0fad954af6ab78c827e Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Fri, 20 Dec 2024 05:03:53 +0000 Subject: [PATCH 807/829] chore: rename {to,of}_With{Upper,Lower}{Set,}_symm_eq (#19996) --- Mathlib/Topology/Order/LowerUpperTopology.lean | 12 ++++++++---- Mathlib/Topology/Order/UpperLowerSetTopology.lean | 12 ++++++++---- 2 files changed, 16 insertions(+), 8 deletions(-) diff --git a/Mathlib/Topology/Order/LowerUpperTopology.lean b/Mathlib/Topology/Order/LowerUpperTopology.lean index 6a77b7876bb67..34ccdbc54eeb5 100644 --- a/Mathlib/Topology/Order/LowerUpperTopology.lean +++ b/Mathlib/Topology/Order/LowerUpperTopology.lean @@ -78,9 +78,11 @@ namespace WithLower /-- `ofLower` is the identity function from the `WithLower` of a type. -/ @[match_pattern] def ofLower : WithLower α ≃ α := Equiv.refl _ -@[simp] lemma to_WithLower_symm_eq : (@toLower α).symm = ofLower := rfl +@[simp] lemma toLower_symm : (@toLower α).symm = ofLower := rfl +@[deprecated (since := "2024-12-16")] alias to_WithLower_symm_eq := toLower_symm -@[simp] lemma of_WithLower_symm_eq : (@ofLower α).symm = toLower := rfl +@[simp] lemma ofLower_symm : (@ofLower α).symm = toLower := rfl +@[deprecated (since := "2024-12-16")] alias of_WithLower_symm_eq := ofLower_symm @[simp] lemma toLower_ofLower (a : WithLower α) : toLower (ofLower a) = a := rfl @@ -122,8 +124,10 @@ namespace WithUpper /-- `ofUpper` is the identity function from the `WithUpper` of a type. -/ @[match_pattern] def ofUpper : WithUpper α ≃ α := Equiv.refl _ -@[simp] lemma to_WithUpper_symm_eq {α} : (@toUpper α).symm = ofUpper := rfl -@[simp] lemma of_WithUpper_symm_eq : (@ofUpper α).symm = toUpper := rfl +@[simp] lemma toUpper_symm {α} : (@toUpper α).symm = ofUpper := rfl +@[deprecated (since := "2024-12-16")] alias to_WithUpper_symm_eq := toUpper_symm +@[simp] lemma ofUpper_symm : (@ofUpper α).symm = toUpper := rfl +@[deprecated (since := "2024-12-16")] alias of_WithUpper_symm_eq := ofUpper_symm @[simp] lemma toUpper_ofUpper (a : WithUpper α) : toUpper (ofUpper a) = a := rfl @[simp] lemma ofUpper_toUpper (a : α) : ofUpper (toUpper a) = a := rfl lemma toUpper_inj {a b : α} : toUpper a = toUpper b ↔ a = b := Iff.rfl diff --git a/Mathlib/Topology/Order/UpperLowerSetTopology.lean b/Mathlib/Topology/Order/UpperLowerSetTopology.lean index 8c2d9de64d0b4..19ca3cbaf5fcf 100644 --- a/Mathlib/Topology/Order/UpperLowerSetTopology.lean +++ b/Mathlib/Topology/Order/UpperLowerSetTopology.lean @@ -80,8 +80,10 @@ namespace WithUpperSet /-- `ofUpperSet` is the identity function from the `WithUpperSet` of a type. -/ @[match_pattern] def ofUpperSet : WithUpperSet α ≃ α := Equiv.refl _ -@[simp] lemma to_WithUpperSet_symm_eq : (@toUpperSet α).symm = ofUpperSet := rfl -@[simp] lemma of_WithUpperSet_symm_eq : (@ofUpperSet α).symm = toUpperSet := rfl +@[simp] lemma toUpperSet_symm : (@toUpperSet α).symm = ofUpperSet := rfl +@[deprecated (since := "2024-10-10")] alias to_WithUpperSet_symm_eq := toUpperSet_symm +@[simp] lemma ofUpperSet_symm : (@ofUpperSet α).symm = toUpperSet := rfl +@[deprecated (since := "2024-10-10")] alias of_WithUpperSet_symm_eq := ofUpperSet_symm @[simp] lemma toUpperSet_ofUpperSet (a : WithUpperSet α) : toUpperSet (ofUpperSet a) = a := rfl @[simp] lemma ofUpperSet_toUpperSet (a : α) : ofUpperSet (toUpperSet a) = a := rfl lemma toUpperSet_inj {a b : α} : toUpperSet a = toUpperSet b ↔ a = b := Iff.rfl @@ -126,8 +128,10 @@ namespace WithLowerSet /-- `ofLowerSet` is the identity function from the `WithLowerSet` of a type. -/ @[match_pattern] def ofLowerSet : WithLowerSet α ≃ α := Equiv.refl _ -@[simp] lemma to_WithLowerSet_symm_eq : (@toLowerSet α).symm = ofLowerSet := rfl -@[simp] lemma of_WithLowerSet_symm_eq : (@ofLowerSet α).symm = toLowerSet := rfl +@[simp] lemma toLowerSet_symm : (@toLowerSet α).symm = ofLowerSet := rfl +@[deprecated (since := "2024-10-10")] alias to_WithLowerSet_symm_eq := toLowerSet_symm +@[simp] lemma ofLowerSet_symm : (@ofLowerSet α).symm = toLowerSet := rfl +@[deprecated (since := "2024-10-10")] alias of_WithLowerSet_symm_eq := ofLowerSet_symm @[simp] lemma toLowerSet_ofLowerSet (a : WithLowerSet α) : toLowerSet (ofLowerSet a) = a := rfl @[simp] lemma ofLowerSet_toLowerSet (a : α) : ofLowerSet (toLowerSet a) = a := rfl lemma toLowerSet_inj {a b : α} : toLowerSet a = toLowerSet b ↔ a = b := Iff.rfl From bee1d7898f11adcfd66bcd56882b136c670acc76 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Fri, 20 Dec 2024 05:03:54 +0000 Subject: [PATCH 808/829] chore: rename ZMod.IsUnit_cast_of_dvd (#19997) --- Mathlib/Data/ZMod/Units.lean | 3 ++- Mathlib/NumberTheory/DirichletCharacter/Basic.lean | 2 +- 2 files changed, 3 insertions(+), 2 deletions(-) diff --git a/Mathlib/Data/ZMod/Units.lean b/Mathlib/Data/ZMod/Units.lean index 3d4def6281661..a39984a54cdcf 100644 --- a/Mathlib/Data/ZMod/Units.lean +++ b/Mathlib/Data/ZMod/Units.lean @@ -32,8 +32,9 @@ lemma unitsMap_comp {d : ℕ} (hm : n ∣ m) (hd : m ∣ d) : lemma unitsMap_self (n : ℕ) : unitsMap (dvd_refl n) = MonoidHom.id _ := by simp [unitsMap, castHom_self] -lemma IsUnit_cast_of_dvd (hm : n ∣ m) (a : Units (ZMod m)) : IsUnit (cast (a : ZMod m) : ZMod n) := +lemma isUnit_cast_of_dvd (hm : n ∣ m) (a : Units (ZMod m)) : IsUnit (cast (a : ZMod m) : ZMod n) := Units.isUnit (unitsMap hm a) +@[deprecated (since := "2024-12-16")] alias IsUnit_cast_of_dvd := isUnit_cast_of_dvd theorem unitsMap_surjective [hm : NeZero m] (h : n ∣ m) : Function.Surjective (unitsMap h) := by diff --git a/Mathlib/NumberTheory/DirichletCharacter/Basic.lean b/Mathlib/NumberTheory/DirichletCharacter/Basic.lean index 8780485d9254f..6e8845121d963 100644 --- a/Mathlib/NumberTheory/DirichletCharacter/Basic.lean +++ b/Mathlib/NumberTheory/DirichletCharacter/Basic.lean @@ -93,7 +93,7 @@ lemma changeLevel_eq_cast_of_dvd {m : ℕ} (hm : n ∣ m) (a : Units (ZMod m)) : (changeLevel hm χ) a = χ (ZMod.cast (a : ZMod m)) := by set_option tactic.skipAssignedInstances false in simpa [changeLevel_def, Function.comp_apply, MonoidHom.coe_comp] using - toUnitHom_eq_char' _ <| ZMod.IsUnit_cast_of_dvd hm a + toUnitHom_eq_char' _ <| ZMod.isUnit_cast_of_dvd hm a /-- `χ` of level `n` factors through a Dirichlet character `χ₀` of level `d` if `d ∣ n` and `χ₀ = χ ∘ (ZMod n → ZMod d)`. -/ From a83bee08cf60822dd1e1d6d3ff4081746bcbb4dd Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Fri, 20 Dec 2024 05:03:55 +0000 Subject: [PATCH 809/829] chore: typo in reference (#19998) --- .../LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean | 4 ++-- Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean | 4 ++-- docs/references.bib | 2 +- 3 files changed, 5 insertions(+), 5 deletions(-) diff --git a/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean b/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean index 0a28708c322c3..d291b7fbb148d 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean @@ -31,8 +31,8 @@ Weyl group. * `rootForm_self_non_neg`: `RootForm` is positive semidefinite. ## References: - * [N. Bourbaki, *Lie groups and {L}ie algebras. {C}hapters 4--6*][bourbaki1968] - * [M. Demazure, *SGA III, Expos\'{e} XXI, Don\'{e}es Radicielles*][demazure1970] + * [N. Bourbaki, *Lie groups and Lie algebras. Chapters 4--6*][bourbaki1968] + * [M. Demazure, *SGA III, Exposé XXI, Données Radicielles*][demazure1970] ## TODO (possibly in other files) * Weyl-invariance diff --git a/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean b/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean index 82b70c781f91a..4157651396831 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean @@ -35,8 +35,8 @@ Weyl group. non-degenerate if the coefficients are a field and the pairing is crystallographic. ## References: - * [N. Bourbaki, *Lie groups and {L}ie algebras. {C}hapters 4--6*][bourbaki1968] - * [M. Demazure, *SGA III, Expos\'{e} XXI, Don\'{e}es Radicielles*][demazure1970] + * [N. Bourbaki, *Lie groups and Lie algebras. Chapters 4--6*][bourbaki1968] + * [M. Demazure, *SGA III, Exposé XXI, Données Radicielles*][demazure1970] ## Todo * Weyl-invariance of `RootForm` and `CorootForm` diff --git a/docs/references.bib b/docs/references.bib index e42ea24949028..3114ae79ad4af 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -1146,7 +1146,7 @@ @InProceedings{ deligne_formulaire @InProceedings{ demazure1970, author = {Michel Demazure}, editor = {M. Demazure, A. Grothendieck}, - title = {Expos\'{e} XXI, Don\'{e}es Radicielles}, + title = {Expos\'{e} XXI, Donn\'{e}es Radicielles}, booktitle = {S\'{e}minaire de G\'{e}ométrie Alg\'{e}brique du Bois Marie - 1962-64 - Sch\'{e}mas en groupes - (SGA 3) - vol. 3, Structure des Sch\'{e}mas en Groupes Reductifs}, From 0e8009b2f95555cae8624ee94e252203d24982b9 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Fri, 20 Dec 2024 05:03:56 +0000 Subject: [PATCH 810/829] chore: simplify YoungDiagram.coe_bot and address porting note (#20000) --- Mathlib/Combinatorics/Young/YoungDiagram.lean | 13 ++++--------- 1 file changed, 4 insertions(+), 9 deletions(-) diff --git a/Mathlib/Combinatorics/Young/YoungDiagram.lean b/Mathlib/Combinatorics/Young/YoungDiagram.lean index 462a840d1e220..b4ca1dc27ecfb 100644 --- a/Mathlib/Combinatorics/Young/YoungDiagram.lean +++ b/Mathlib/Combinatorics/Young/YoungDiagram.lean @@ -152,19 +152,14 @@ instance : OrderBot YoungDiagram where theorem cells_bot : (⊥ : YoungDiagram).cells = ∅ := rfl --- Porting note: removed `↑`, added `.cells` and changed proof -@[norm_cast] -theorem coe_bot : (⊥ : YoungDiagram).cells = (∅ : Set (ℕ × ℕ)) := by - refine Set.eq_of_subset_of_subset ?_ ?_ - · intros x h - simp? [mem_mk, Finset.coe_empty, Set.mem_empty_iff_false] at h says - simp only [cells_bot, Finset.coe_empty, Set.mem_empty_iff_false] at h - · simp only [cells_bot, Finset.coe_empty, Set.empty_subset] - @[simp] theorem not_mem_bot (x : ℕ × ℕ) : x ∉ (⊥ : YoungDiagram) := Finset.not_mem_empty x +@[norm_cast] +theorem coe_bot : (⊥ : YoungDiagram) = (∅ : Set (ℕ × ℕ)) := by + ext; simp + instance : Inhabited YoungDiagram := ⟨⊥⟩ From f4eac440fe2bb4d4b6e73e5f186341d7e3518b11 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 20 Dec 2024 05:03:57 +0000 Subject: [PATCH 811/829] =?UTF-8?q?feat:=20`1=20=E2=89=A4=20a=20^=20n=20?= =?UTF-8?q?=E2=86=94=200=20=E2=89=A4=20n`=20when=20`1=20<=20a`=20(#20032)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit and other basic `zpow` lemmas --- .../Order/GroupWithZero/Unbundled.lean | 26 ++++++++++++++++++- 1 file changed, 25 insertions(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean index a9286e9934add..b9105af63bd8c 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean @@ -1618,6 +1618,30 @@ lemma zpow_le_zpow_iff_right_of_lt_one₀ (ha₀ : 0 < a) (ha₁ : a < 1) : lemma zpow_lt_zpow_iff_right_of_lt_one₀ (ha₀ : 0 < a) (ha₁ : a < 1) : a ^ m < a ^ n ↔ n < m := (zpow_right_strictAnti₀ ha₀ ha₁).lt_iff_lt +@[simp] lemma one_le_zpow_iff_right₀ (ha : 1 < a) : 1 ≤ a ^ n ↔ 0 ≤ n := by + simp [← zpow_le_zpow_iff_right₀ ha] + +@[simp] lemma one_lt_zpow_iff_right₀ (ha : 1 < a) : 1 < a ^ n ↔ 0 < n := by + simp [← zpow_lt_zpow_iff_right₀ ha] + +@[simp] lemma one_le_zpow_iff_right_of_lt_one₀ (ha₀ : 0 < a) (ha₁ : a < 1) : 1 ≤ a ^ n ↔ n ≤ 0 := by + simp [← zpow_le_zpow_iff_right_of_lt_one₀ ha₀ ha₁] + +@[simp] lemma one_lt_zpow_iff_right_of_lt_one₀ (ha₀ : 0 < a) (ha₁ : a < 1) : 1 < a ^ n ↔ n < 0 := by + simp [← zpow_lt_zpow_iff_right_of_lt_one₀ ha₀ ha₁] + +@[simp] lemma zpow_le_one_iff_right₀ (ha : 1 < a) : a ^ n ≤ 1 ↔ n ≤ 0 := by + simp [← zpow_le_zpow_iff_right₀ ha] + +@[simp] lemma zpow_lt_one_iff_right₀ (ha : 1 < a) : a ^ n < 1 ↔ n < 0 := by + simp [← zpow_lt_zpow_iff_right₀ ha] + +@[simp] lemma zpow_le_one_iff_right_of_lt_one₀ (ha₀ : 0 < a) (ha₁ : a < 1) : a ^ n ≤ 1 ↔ 0 ≤ n := by + simp [← zpow_le_zpow_iff_right_of_lt_one₀ ha₀ ha₁] + +@[simp] lemma zpow_lt_one_iff_right_of_lt_one₀ (ha₀ : 0 < a) (ha₁ : a < 1) : a ^ n < 1 ↔ 0 < n := by + simp [← zpow_lt_zpow_iff_right_of_lt_one₀ ha₀ ha₁] + end PosMulStrictMono section MulPosStrictMono @@ -1882,4 +1906,4 @@ lemma div_lt_comm₀ (hb : 0 < b) (hc : 0 < c) : a / b < c ↔ a / c < b := by end PosMulStrictMono end CommGroupWithZero -set_option linter.style.longFile 1900 +set_option linter.style.longFile 2000 From 5999d673061ac457e356cfb1e3cdc24da41e8b14 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 20 Dec 2024 05:03:59 +0000 Subject: [PATCH 812/829] chore(Multiset): make all arguments to `inter_le_left`, etc... implicit (#20045) This follows what we already do for `Set`, `Finset`, general lattices --- Mathlib/Data/Multiset/Basic.lean | 78 +++++++++++++--------------- Mathlib/Data/Multiset/FinsetOps.lean | 4 +- Mathlib/Data/Multiset/Nodup.lean | 9 ++-- 3 files changed, 41 insertions(+), 50 deletions(-) diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index 7eabbce0e4d7e..f7425d084e924 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -1543,10 +1543,10 @@ instance : Union (Multiset α) := theorem union_def (s t : Multiset α) : s ∪ t = s - t + t := rfl -theorem le_union_left (s t : Multiset α) : s ≤ s ∪ t := +theorem le_union_left {s t : Multiset α} : s ≤ s ∪ t := le_tsub_add -theorem le_union_right (s t : Multiset α) : t ≤ s ∪ t := +theorem le_union_right {s t : Multiset α} : t ≤ s ∪ t := le_add_left _ _ theorem eq_union_left : t ≤ s → s ∪ t = s := @@ -1563,7 +1563,7 @@ theorem union_le (h₁ : s ≤ u) (h₂ : t ≤ u) : s ∪ t ≤ u := by @[simp] theorem mem_union : a ∈ s ∪ t ↔ a ∈ s ∨ a ∈ t := ⟨fun h => (mem_add.1 h).imp_left (mem_of_le <| Multiset.sub_le_self _ _), - (Or.elim · (mem_of_le <| le_union_left _ _) (mem_of_le <| le_union_right _ _))⟩ + (Or.elim · (mem_of_le le_union_left) (mem_of_le le_union_right))⟩ @[simp] theorem map_union [DecidableEq β] {f : α → β} (finj : Function.Injective f) {s t : Multiset α} : @@ -1605,12 +1605,15 @@ theorem cons_inter_of_pos {a} (s : Multiset α) {t} : a ∈ t → (a ::ₘ s) theorem cons_inter_of_neg {a} (s : Multiset α) {t} : a ∉ t → (a ::ₘ s) ∩ t = s ∩ t := Quotient.inductionOn₂ s t fun _l₁ _l₂ h => congr_arg ofList <| cons_bagInter_of_neg _ h -theorem inter_le_left (s t : Multiset α) : s ∩ t ≤ s := +theorem inter_le_left {s t : Multiset α} : s ∩ t ≤ s := Quotient.inductionOn₂ s t fun _l₁ _l₂ => (bagInter_sublist_left _ _).subperm -theorem inter_le_right (s : Multiset α) : ∀ t, s ∩ t ≤ t := - Multiset.induction_on s (fun t => (zero_inter t).symm ▸ zero_le _) fun a s IH t => - if h : a ∈ t then by simpa [h] using cons_le_cons a (IH (t.erase a)) else by simp [h, IH] +theorem inter_le_right {s t : Multiset α} : s ∩ t ≤ t := by + induction' s using Multiset.induction_on with a s IH generalizing t + · exact (zero_inter t).symm ▸ zero_le _ + by_cases h : a ∈ t + · simpa [h] using cons_le_cons a (IH (t := t.erase a)) + · simp [h, IH] theorem le_inter (h₁ : s ≤ t) (h₂ : s ≤ u) : s ≤ t ∩ u := by revert s u; refine @(Multiset.induction_on t ?_ fun a t IH => ?_) <;> intros s u h₁ h₂ @@ -1623,18 +1626,18 @@ theorem le_inter (h₁ : s ≤ t) (h₂ : s ≤ u) : s ≤ t ∩ u := by @[simp] theorem mem_inter : a ∈ s ∩ t ↔ a ∈ s ∧ a ∈ t := - ⟨fun h => ⟨mem_of_le (inter_le_left _ _) h, mem_of_le (inter_le_right _ _) h⟩, fun ⟨h₁, h₂⟩ => by + ⟨fun h => ⟨mem_of_le inter_le_left h, mem_of_le inter_le_right h⟩, fun ⟨h₁, h₂⟩ => by rw [← cons_erase h₁, cons_inter_of_pos _ h₂]; apply mem_cons_self⟩ -instance : Lattice (Multiset α) := - { sup := (· ∪ ·) - sup_le := @union_le _ _ - le_sup_left := le_union_left - le_sup_right := le_union_right - inf := (· ∩ ·) - le_inf := @le_inter _ _ - inf_le_left := inter_le_left - inf_le_right := inter_le_right } +instance : Lattice (Multiset α) where + sup := (· ∪ ·) + sup_le _ _ _ := union_le + le_sup_left _ _ := le_union_left + le_sup_right _ _ := le_union_right + inf := (· ∩ ·) + le_inf _ _ _ := le_inter + inf_le_left _ _ := inter_le_left + inf_le_right _ _ := inter_le_right @[simp] theorem sup_eq_union (s t : Multiset α) : s ⊔ t = s ∪ t := @@ -1676,19 +1679,13 @@ theorem cons_union_distrib (a : α) (s t : Multiset α) : a ::ₘ (s ∪ t) = a simpa using add_union_distrib (a ::ₘ 0) s t theorem inter_add_distrib (s t u : Multiset α) : s ∩ t + u = (s + u) ∩ (t + u) := by - by_contra h - cases' - lt_iff_cons_le.1 - (lt_of_le_of_ne - (le_inter (add_le_add_right (inter_le_left s t) u) - (add_le_add_right (inter_le_right s t) u)) - h) with - a hl - rw [← cons_add] at hl - exact - not_le_of_lt (lt_cons_self (s ∩ t) a) - (le_inter (le_of_add_le_add_right (le_trans hl (inter_le_left _ _))) - (le_of_add_le_add_right (le_trans hl (inter_le_right _ _)))) + by_contra! h + obtain ⟨a, ha⟩ := lt_iff_cons_le.1 <| h.lt_of_le <| le_inter + (add_le_add_right inter_le_left _) (add_le_add_right inter_le_right _) + rw [← cons_add] at ha + exact (lt_cons_self (s ∩ t) a).not_le <| le_inter + (le_of_add_le_add_right (ha.trans inter_le_left)) + (le_of_add_le_add_right (ha.trans inter_le_right)) theorem add_inter_distrib (s t u : Multiset α) : s + t ∩ u = (s + t) ∩ (s + u) := by rw [add_comm, inter_add_distrib, add_comm s, add_comm s] @@ -1696,16 +1693,16 @@ theorem add_inter_distrib (s t u : Multiset α) : s + t ∩ u = (s + t) ∩ (s + theorem cons_inter_distrib (a : α) (s t : Multiset α) : a ::ₘ s ∩ t = (a ::ₘ s) ∩ (a ::ₘ t) := by simp -theorem union_add_inter (s t : Multiset α) : s ∪ t + s ∩ t = s + t := by +lemma union_add_inter (s t : Multiset α) : s ∪ t + s ∩ t = s + t := by apply _root_.le_antisymm · rw [union_add_distrib] - refine union_le (add_le_add_left (inter_le_right _ _) _) ?_ + refine union_le (add_le_add_left inter_le_right _) ?_ rw [add_comm] - exact add_le_add_right (inter_le_left _ _) _ + exact add_le_add_right inter_le_left _ · rw [add_comm, add_inter_distrib] - refine le_inter (add_le_add_right (le_union_right _ _) _) ?_ + refine le_inter (add_le_add_right le_union_right _) ?_ rw [add_comm] - exact add_le_add_right (le_union_left _ _) _ + exact add_le_add_right le_union_left _ theorem sub_add_inter (s t : Multiset α) : s - t + s ∩ t = s := by rw [inter_comm] @@ -1716,7 +1713,7 @@ theorem sub_add_inter (s t : Multiset α) : s - t + s ∩ t = s := by theorem sub_inter (s t : Multiset α) : s - s ∩ t = s - t := add_right_cancel (b := s ∩ t) <| by - rw [sub_add_inter s t, tsub_add_cancel_of_le (inter_le_left s t)] + rw [sub_add_inter s t, tsub_add_cancel_of_le inter_le_left] end @@ -1850,12 +1847,9 @@ theorem filter_union [DecidableEq α] (s t : Multiset α) : @[simp] theorem filter_inter [DecidableEq α] (s t : Multiset α) : filter p (s ∩ t) = filter p s ∩ filter p t := - le_antisymm - (le_inter (filter_le_filter _ <| inter_le_left _ _) - (filter_le_filter _ <| inter_le_right _ _)) <| - le_filter.2 - ⟨inf_le_inf (filter_le _ _) (filter_le _ _), fun _a h => - of_mem_filter (mem_of_le (inter_le_left _ _) h)⟩ + le_antisymm (le_inter (filter_le_filter _ inter_le_left) (filter_le_filter _ inter_le_right)) <| + le_filter.2 ⟨inf_le_inf (filter_le _ _) (filter_le _ _), fun _a h => + of_mem_filter (mem_of_le inter_le_left h)⟩ @[simp] theorem filter_filter (q) [DecidablePred q] (s : Multiset α) : diff --git a/Mathlib/Data/Multiset/FinsetOps.lean b/Mathlib/Data/Multiset/FinsetOps.lean index 2cb3a982b4c23..0a99e9258eaf2 100644 --- a/Mathlib/Data/Multiset/FinsetOps.lean +++ b/Mathlib/Data/Multiset/FinsetOps.lean @@ -164,7 +164,7 @@ theorem le_ndunion_left {s} (t : Multiset α) (d : Nodup s) : s ≤ ndunion s t (le_iff_subset d).2 <| subset_ndunion_left _ _ theorem ndunion_le_union (s t : Multiset α) : ndunion s t ≤ s ∪ t := - ndunion_le.2 ⟨subset_of_le (le_union_left _ _), le_union_right _ _⟩ + ndunion_le.2 ⟨subset_of_le le_union_left, le_union_right⟩ theorem Nodup.ndunion (s : Multiset α) {t : Multiset α} : Nodup t → Nodup (ndunion s t) := Quot.induction_on₂ s t fun _ _ => List.Nodup.union _ @@ -236,7 +236,7 @@ theorem ndinter_le_right {s} (t : Multiset α) (d : Nodup s) : ndinter s t ≤ t (le_iff_subset <| d.ndinter _).2 <| ndinter_subset_right _ _ theorem inter_le_ndinter (s t : Multiset α) : s ∩ t ≤ ndinter s t := - le_ndinter.2 ⟨inter_le_left _ _, subset_of_le <| inter_le_right _ _⟩ + le_ndinter.2 ⟨inter_le_left, subset_of_le inter_le_right⟩ @[simp] theorem ndinter_eq_inter {s t : Multiset α} (d : Nodup s) : ndinter s t = s ∩ t := diff --git a/Mathlib/Data/Multiset/Nodup.lean b/Mathlib/Data/Multiset/Nodup.lean index f3fc2f99651f9..87bbfbd63a93a 100644 --- a/Mathlib/Data/Multiset/Nodup.lean +++ b/Mathlib/Data/Multiset/Nodup.lean @@ -167,15 +167,12 @@ protected theorem Nodup.filterMap (f : α → Option β) (H : ∀ a a' b, b ∈ theorem nodup_range (n : ℕ) : Nodup (range n) := List.nodup_range _ -theorem Nodup.inter_left [DecidableEq α] (t) : Nodup s → Nodup (s ∩ t) := - nodup_of_le <| inter_le_left _ _ - -theorem Nodup.inter_right [DecidableEq α] (s) : Nodup t → Nodup (s ∩ t) := - nodup_of_le <| inter_le_right _ _ +lemma Nodup.inter_left [DecidableEq α] (t) : Nodup s → Nodup (s ∩ t) := nodup_of_le inter_le_left +lemma Nodup.inter_right [DecidableEq α] (s) : Nodup t → Nodup (s ∩ t) := nodup_of_le inter_le_right @[simp] theorem nodup_union [DecidableEq α] {s t : Multiset α} : Nodup (s ∪ t) ↔ Nodup s ∧ Nodup t := - ⟨fun h => ⟨nodup_of_le (le_union_left _ _) h, nodup_of_le (le_union_right _ _) h⟩, fun ⟨h₁, h₂⟩ => + ⟨fun h => ⟨nodup_of_le le_union_left h, nodup_of_le le_union_right h⟩, fun ⟨h₁, h₂⟩ => nodup_iff_count_le_one.2 fun a => by rw [count_union] exact max_le (nodup_iff_count_le_one.1 h₁ a) (nodup_iff_count_le_one.1 h₂ a)⟩ From 97ba55d9cd61f902b13bda9cb7c1d9b96befe830 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Fri, 20 Dec 2024 05:04:00 +0000 Subject: [PATCH 813/829] chore(Analysis/InnerProductSpace): split large file (#20046) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Split up `Analysis.InnerProductSpace.Basic` as follows: - `Defs`: the definition and the `InnerProductSpace.Core` stuff - `Basic`: general lemmas - `Completion`: inner products on completions and separation quotients - `LinearMap`: interaction of linear maps with inner products - `Orthonormal`: orthonormal families of vectors - `Subspace`: subspaces of inner product spaces (including orthogonal families) No code changes, other than deleting one copy of an instance which was declared twice. The file `Defs` uses `ℂ` in a notation definition, which is invisible to the tree-shake script, so it needs to have an exception in `noshake.json`. --- Mathlib.lean | 5 + Mathlib/Analysis/InnerProductSpace/Basic.lean | 1543 +---------------- .../InnerProductSpace/Completion.lean | 119 ++ .../InnerProductSpace/ConformalLinearMap.lean | 2 +- Mathlib/Analysis/InnerProductSpace/Defs.lean | 502 ++++++ .../Analysis/InnerProductSpace/LinearMap.lean | 308 ++++ .../InnerProductSpace/Orthogonal.lean | 2 +- .../InnerProductSpace/Orthonormal.lean | 444 +++++ .../Analysis/InnerProductSpace/Subspace.lean | 275 +++ .../Analysis/InnerProductSpace/Symmetric.lean | 2 +- .../Euclidean/Angle/Unoriented/Basic.lean | 2 +- Mathlib/MeasureTheory/Function/L2Space.lean | 2 +- scripts/noshake.json | 3 +- 13 files changed, 1668 insertions(+), 1541 deletions(-) create mode 100644 Mathlib/Analysis/InnerProductSpace/Completion.lean create mode 100644 Mathlib/Analysis/InnerProductSpace/Defs.lean create mode 100644 Mathlib/Analysis/InnerProductSpace/LinearMap.lean create mode 100644 Mathlib/Analysis/InnerProductSpace/Orthonormal.lean create mode 100644 Mathlib/Analysis/InnerProductSpace/Subspace.lean diff --git a/Mathlib.lean b/Mathlib.lean index 88d19be469fa5..5f0ed064cabd5 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1330,18 +1330,22 @@ import Mathlib.Analysis.Hofer import Mathlib.Analysis.InnerProductSpace.Adjoint import Mathlib.Analysis.InnerProductSpace.Basic import Mathlib.Analysis.InnerProductSpace.Calculus +import Mathlib.Analysis.InnerProductSpace.Completion import Mathlib.Analysis.InnerProductSpace.ConformalLinearMap +import Mathlib.Analysis.InnerProductSpace.Defs import Mathlib.Analysis.InnerProductSpace.Dual import Mathlib.Analysis.InnerProductSpace.EuclideanDist import Mathlib.Analysis.InnerProductSpace.GramSchmidtOrtho import Mathlib.Analysis.InnerProductSpace.JointEigenspace import Mathlib.Analysis.InnerProductSpace.LaxMilgram +import Mathlib.Analysis.InnerProductSpace.LinearMap import Mathlib.Analysis.InnerProductSpace.LinearPMap import Mathlib.Analysis.InnerProductSpace.MeanErgodic import Mathlib.Analysis.InnerProductSpace.NormPow import Mathlib.Analysis.InnerProductSpace.OfNorm import Mathlib.Analysis.InnerProductSpace.Orientation import Mathlib.Analysis.InnerProductSpace.Orthogonal +import Mathlib.Analysis.InnerProductSpace.Orthonormal import Mathlib.Analysis.InnerProductSpace.PiL2 import Mathlib.Analysis.InnerProductSpace.Positive import Mathlib.Analysis.InnerProductSpace.ProdL2 @@ -1350,6 +1354,7 @@ import Mathlib.Analysis.InnerProductSpace.Rayleigh import Mathlib.Analysis.InnerProductSpace.Semisimple import Mathlib.Analysis.InnerProductSpace.Spectrum import Mathlib.Analysis.InnerProductSpace.StarOrder +import Mathlib.Analysis.InnerProductSpace.Subspace import Mathlib.Analysis.InnerProductSpace.Symmetric import Mathlib.Analysis.InnerProductSpace.TwoDim import Mathlib.Analysis.InnerProductSpace.WeakOperatorTopology diff --git a/Mathlib/Analysis/InnerProductSpace/Basic.lean b/Mathlib/Analysis/InnerProductSpace/Basic.lean index 6f06b27cb078c..377de45952c82 100644 --- a/Mathlib/Analysis/InnerProductSpace/Basic.lean +++ b/Mathlib/Analysis/InnerProductSpace/Basic.lean @@ -3,511 +3,41 @@ Copyright (c) 2019 Zhouhang Zhou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Zhouhang Zhou, Sébastien Gouëzel, Frédéric Dupuis -/ -import Mathlib.Algebra.DirectSum.Module -import Mathlib.Algebra.Module.LinearMap.Basic -import Mathlib.Algebra.QuadraticDiscriminant + import Mathlib.Analysis.Complex.Basic import Mathlib.Analysis.Convex.Uniform -import Mathlib.Analysis.Normed.Module.Completion +import Mathlib.Analysis.InnerProductSpace.Defs import Mathlib.Analysis.Normed.Operator.BoundedLinearMaps /-! -# Inner product space - -This file defines inner product spaces and proves the basic properties. We do not formally -define Hilbert spaces, but they can be obtained using the set of assumptions -`[NormedAddCommGroup E] [InnerProductSpace 𝕜 E] [CompleteSpace E]`. - -An inner product space is a vector space endowed with an inner product. It generalizes the notion of -dot product in `ℝ^n` and provides the means of defining the length of a vector and the angle between -two vectors. In particular vectors `x` and `y` are orthogonal if their inner product equals zero. -We define both the real and complex cases at the same time using the `RCLike` typeclass. +# Properties of inner product spaces -This file proves general results on inner product spaces. For the specific construction of an inner -product structure on `n → 𝕜` for `𝕜 = ℝ` or `ℂ`, see `EuclideanSpace` in -`Analysis.InnerProductSpace.PiL2`. +This file proves many basic properties of inner product spaces (real or complex). ## Main results -- We define the class `InnerProductSpace 𝕜 E` extending `NormedSpace 𝕜 E` with a number of basic - properties, most notably the Cauchy-Schwarz inequality. Here `𝕜` is understood to be either `ℝ` - or `ℂ`, through the `RCLike` typeclass. +- `inner_mul_inner_self_le`: the Cauchy-Schwartz inequality (one of many variants). +- `norm_inner_eq_norm_iff`: the equality criteion in the Cauchy-Schwartz inequality (also in many + variants). +- `inner_eq_sum_norm_sq_div_four`: the polarization identity. - We show that the inner product is continuous, `continuous_inner`, and bundle it as the continuous sesquilinear map `innerSL` (see also `innerₛₗ` for the non-continuous version). -- We define `Orthonormal`, a predicate on a function `v : ι → E`, and prove the existence of a - maximal orthonormal set, `exists_maximal_orthonormal`. Bessel's inequality, - `Orthonormal.tsum_inner_products_le`, states that given an orthonormal set `v` and a vector `x`, - the sum of the norm-squares of the inner products `⟪v i, x⟫` is no more than the norm-square of - `x`. For the existence of orthonormal bases, Hilbert bases, etc., see the file - `Analysis.InnerProductSpace.projection`. - -## Notation - -We globally denote the real and complex inner products by `⟪·, ·⟫_ℝ` and `⟪·, ·⟫_ℂ` respectively. -We also provide two notation namespaces: `RealInnerProductSpace`, `ComplexInnerProductSpace`, -which respectively introduce the plain notation `⟪·, ·⟫` for the real and complex inner product. - -## Implementation notes - -We choose the convention that inner products are conjugate linear in the first argument and linear -in the second. ## Tags inner product space, Hilbert space, norm -## References -* [Clément & Martin, *The Lax-Milgram Theorem. A detailed proof to be formalized in Coq*] -* [Clément & Martin, *A Coq formal proof of the Lax–Milgram theorem*] - -The Coq code is available at the following address: -/ noncomputable section -open RCLike Real Filter - -open Topology ComplexConjugate Finsupp +open RCLike Real Filter Topology ComplexConjugate Finsupp open LinearMap (BilinForm) variable {𝕜 E F : Type*} [RCLike 𝕜] -/-- Syntactic typeclass for types endowed with an inner product -/ -class Inner (𝕜 E : Type*) where - /-- The inner product function. -/ - inner : E → E → 𝕜 - -export Inner (inner) - -/-- The inner product with values in `𝕜`. -/ -scoped[InnerProductSpace] notation3:max "⟪" x ", " y "⟫_" 𝕜:max => @inner 𝕜 _ _ x y - -section Notations - -/-- The inner product with values in `ℝ`. -/ -scoped[RealInnerProductSpace] notation "⟪" x ", " y "⟫" => @inner ℝ _ _ x y - -/-- The inner product with values in `ℂ`. -/ -scoped[ComplexInnerProductSpace] notation "⟪" x ", " y "⟫" => @inner ℂ _ _ x y - -end Notations - -/-- A (pre) inner product space is a vector space with an additional operation called inner product. -The (semi)norm could be derived from the inner product, instead we require the existence of a -seminorm and the fact that `‖x‖^2 = re ⟪x, x⟫` to be able to put instances on `𝕂` or product spaces. - -Note that `NormedSpace` does not assume that `‖x‖=0` implies `x=0` (it is rather a seminorm). - -To construct a seminorm from an inner product, see `PreInnerProductSpace.ofCore`. --/ -class InnerProductSpace (𝕜 : Type*) (E : Type*) [RCLike 𝕜] [SeminormedAddCommGroup E] extends - NormedSpace 𝕜 E, Inner 𝕜 E where - /-- The inner product induces the norm. -/ - norm_sq_eq_inner : ∀ x : E, ‖x‖ ^ 2 = re (inner x x) - /-- The inner product is *hermitian*, taking the `conj` swaps the arguments. -/ - conj_symm : ∀ x y, conj (inner y x) = inner x y - /-- The inner product is additive in the first coordinate. -/ - add_left : ∀ x y z, inner (x + y) z = inner x z + inner y z - /-- The inner product is conjugate linear in the first coordinate. -/ - smul_left : ∀ x y r, inner (r • x) y = conj r * inner x y - -/-! -### Constructing a normed space structure from an inner product - -In the definition of an inner product space, we require the existence of a norm, which is equal -(but maybe not defeq) to the square root of the scalar product. This makes it possible to put -an inner product space structure on spaces with a preexisting norm (for instance `ℝ`), with good -properties. However, sometimes, one would like to define the norm starting only from a well-behaved -scalar product. This is what we implement in this paragraph, starting from a structure -`InnerProductSpace.Core` stating that we have a nice scalar product. - -Our goal here is not to develop a whole theory with all the supporting API, as this will be done -below for `InnerProductSpace`. Instead, we implement the bare minimum to go as directly as -possible to the construction of the norm and the proof of the triangular inequality. - -Warning: Do not use this `Core` structure if the space you are interested in already has a norm -instance defined on it, otherwise this will create a second non-defeq norm instance! --/ - -/-- A structure requiring that a scalar product is positive semidefinite and symmetric. -/ -structure PreInnerProductSpace.Core (𝕜 : Type*) (F : Type*) [RCLike 𝕜] [AddCommGroup F] - [Module 𝕜 F] extends Inner 𝕜 F where - /-- The inner product is *hermitian*, taking the `conj` swaps the arguments. -/ - conj_symm : ∀ x y, conj (inner y x) = inner x y - /-- The inner product is positive (semi)definite. -/ - nonneg_re : ∀ x, 0 ≤ re (inner x x) - /-- The inner product is additive in the first coordinate. -/ - add_left : ∀ x y z, inner (x + y) z = inner x z + inner y z - /-- The inner product is conjugate linear in the first coordinate. -/ - smul_left : ∀ x y r, inner (r • x) y = conj r * inner x y - -attribute [class] PreInnerProductSpace.Core - -/-- A structure requiring that a scalar product is positive definite. Some theorems that -require this assumptions are put under section `InnerProductSpace.Core`. -/ --- @[nolint HasNonemptyInstance] porting note: I don't think we have this linter anymore -structure InnerProductSpace.Core (𝕜 : Type*) (F : Type*) [RCLike 𝕜] [AddCommGroup F] - [Module 𝕜 F] extends PreInnerProductSpace.Core 𝕜 F where - /-- The inner product is positive definite. -/ - definite : ∀ x, inner x x = 0 → x = 0 - -/- We set `InnerProductSpace.Core` to be a class as we will use it as such in the construction -of the normed space structure that it produces. However, all the instances we will use will be -local to this proof. -/ -attribute [class] InnerProductSpace.Core - -instance (𝕜 : Type*) (F : Type*) [RCLike 𝕜] [AddCommGroup F] - [Module 𝕜 F] [cd : InnerProductSpace.Core 𝕜 F] : PreInnerProductSpace.Core 𝕜 F where - inner := cd.inner - conj_symm := cd.conj_symm - nonneg_re := cd.nonneg_re - add_left := cd.add_left - smul_left := cd.smul_left - -/-- Define `PreInnerProductSpace.Core` from `PreInnerProductSpace`. Defined to reuse lemmas about -`PreInnerProductSpace.Core` for `PreInnerProductSpace`s. Note that the `Seminorm` instance provided -by `PreInnerProductSpace.Core.norm` is propositionally but not definitionally equal to the original -norm. -/ -def PreInnerProductSpace.toCore [SeminormedAddCommGroup E] [c : InnerProductSpace 𝕜 E] : - PreInnerProductSpace.Core 𝕜 E := - { c with - nonneg_re := fun x => by - rw [← InnerProductSpace.norm_sq_eq_inner] - apply sq_nonneg } - -/-- Define `InnerProductSpace.Core` from `InnerProductSpace`. Defined to reuse lemmas about -`InnerProductSpace.Core` for `InnerProductSpace`s. Note that the `Norm` instance provided by -`InnerProductSpace.Core.norm` is propositionally but not definitionally equal to the original -norm. -/ -def InnerProductSpace.toCore [NormedAddCommGroup E] [c : InnerProductSpace 𝕜 E] : - InnerProductSpace.Core 𝕜 E := - { c with - nonneg_re := fun x => by - rw [← InnerProductSpace.norm_sq_eq_inner] - apply sq_nonneg - definite := fun x hx => - norm_eq_zero.1 <| pow_eq_zero (n := 2) <| by - rw [InnerProductSpace.norm_sq_eq_inner (𝕜 := 𝕜) x, hx, map_zero] } - -namespace InnerProductSpace.Core - -section PreInnerProductSpace.Core - -variable [AddCommGroup F] [Module 𝕜 F] [c : PreInnerProductSpace.Core 𝕜 F] - -local notation "⟪" x ", " y "⟫" => @inner 𝕜 F _ x y - -local notation "normSqK" => @RCLike.normSq 𝕜 _ - -local notation "reK" => @RCLike.re 𝕜 _ - -local notation "ext_iff" => @RCLike.ext_iff 𝕜 _ - -local postfix:90 "†" => starRingEnd _ - -/-- Inner product defined by the `PreInnerProductSpace.Core` structure. We can't reuse -`PreInnerProductSpace.Core.toInner` because it takes `PreInnerProductSpace.Core` as an explicit -argument. -/ -def toPreInner' : Inner 𝕜 F := - c.toInner - -attribute [local instance] toPreInner' - -/-- The norm squared function for `PreInnerProductSpace.Core` structure. -/ -def normSq (x : F) := - reK ⟪x, x⟫ - -local notation "normSqF" => @normSq 𝕜 F _ _ _ _ - -theorem inner_conj_symm (x y : F) : ⟪y, x⟫† = ⟪x, y⟫ := - c.conj_symm x y - -theorem inner_self_nonneg {x : F} : 0 ≤ re ⟪x, x⟫ := - c.nonneg_re _ - -theorem inner_self_im (x : F) : im ⟪x, x⟫ = 0 := by - rw [← @ofReal_inj 𝕜, im_eq_conj_sub] - simp [inner_conj_symm] - -theorem inner_add_left (x y z : F) : ⟪x + y, z⟫ = ⟪x, z⟫ + ⟪y, z⟫ := - c.add_left _ _ _ - -theorem inner_add_right (x y z : F) : ⟪x, y + z⟫ = ⟪x, y⟫ + ⟪x, z⟫ := by - rw [← inner_conj_symm, inner_add_left, RingHom.map_add]; simp only [inner_conj_symm] - -theorem ofReal_normSq_eq_inner_self (x : F) : (normSqF x : 𝕜) = ⟪x, x⟫ := by - rw [ext_iff] - exact ⟨by simp only [ofReal_re]; rfl, by simp only [inner_self_im, ofReal_im]⟩ - -theorem inner_re_symm (x y : F) : re ⟪x, y⟫ = re ⟪y, x⟫ := by rw [← inner_conj_symm, conj_re] - -theorem inner_im_symm (x y : F) : im ⟪x, y⟫ = -im ⟪y, x⟫ := by rw [← inner_conj_symm, conj_im] - -theorem inner_smul_left (x y : F) {r : 𝕜} : ⟪r • x, y⟫ = r† * ⟪x, y⟫ := - c.smul_left _ _ _ - -theorem inner_smul_right (x y : F) {r : 𝕜} : ⟪x, r • y⟫ = r * ⟪x, y⟫ := by - rw [← inner_conj_symm, inner_smul_left] - simp only [conj_conj, inner_conj_symm, RingHom.map_mul] - -theorem inner_zero_left (x : F) : ⟪0, x⟫ = 0 := by - rw [← zero_smul 𝕜 (0 : F), inner_smul_left] - simp only [zero_mul, RingHom.map_zero] - -theorem inner_zero_right (x : F) : ⟪x, 0⟫ = 0 := by - rw [← inner_conj_symm, inner_zero_left]; simp only [RingHom.map_zero] - -theorem inner_self_of_eq_zero {x : F} : x = 0 → ⟪x, x⟫ = 0 := by - rintro rfl - exact inner_zero_left _ - -theorem normSq_eq_zero_of_eq_zero {x : F} : x = 0 → normSqF x = 0 := by - rintro rfl - simp [normSq, inner_self_of_eq_zero] - -theorem ne_zero_of_inner_self_ne_zero {x : F} : ⟪x, x⟫ ≠ 0 → x ≠ 0 := - mt inner_self_of_eq_zero -theorem inner_self_ofReal_re (x : F) : (re ⟪x, x⟫ : 𝕜) = ⟪x, x⟫ := by - norm_num [ext_iff, inner_self_im] - -theorem norm_inner_symm (x y : F) : ‖⟪x, y⟫‖ = ‖⟪y, x⟫‖ := by rw [← inner_conj_symm, norm_conj] - -theorem inner_neg_left (x y : F) : ⟪-x, y⟫ = -⟪x, y⟫ := by - rw [← neg_one_smul 𝕜 x, inner_smul_left] - simp - -theorem inner_neg_right (x y : F) : ⟪x, -y⟫ = -⟪x, y⟫ := by - rw [← inner_conj_symm, inner_neg_left]; simp only [RingHom.map_neg, inner_conj_symm] - -theorem inner_sub_left (x y z : F) : ⟪x - y, z⟫ = ⟪x, z⟫ - ⟪y, z⟫ := by - simp [sub_eq_add_neg, inner_add_left, inner_neg_left] - -theorem inner_sub_right (x y z : F) : ⟪x, y - z⟫ = ⟪x, y⟫ - ⟪x, z⟫ := by - simp [sub_eq_add_neg, inner_add_right, inner_neg_right] - -theorem inner_mul_symm_re_eq_norm (x y : F) : re (⟪x, y⟫ * ⟪y, x⟫) = ‖⟪x, y⟫ * ⟪y, x⟫‖ := by - rw [← inner_conj_symm, mul_comm] - exact re_eq_norm_of_mul_conj (inner y x) - -/-- Expand `inner (x + y) (x + y)` -/ -theorem inner_add_add_self (x y : F) : ⟪x + y, x + y⟫ = ⟪x, x⟫ + ⟪x, y⟫ + ⟪y, x⟫ + ⟪y, y⟫ := by - simp only [inner_add_left, inner_add_right]; ring - --- Expand `inner (x - y) (x - y)` -theorem inner_sub_sub_self (x y : F) : ⟪x - y, x - y⟫ = ⟪x, x⟫ - ⟪x, y⟫ - ⟪y, x⟫ + ⟪y, y⟫ := by - simp only [inner_sub_left, inner_sub_right]; ring - -theorem inner_smul_ofReal_left (x y : F) {t : ℝ} : ⟪(t : 𝕜) • x, y⟫ = ⟪x, y⟫ * t := by - rw [inner_smul_left, conj_ofReal, mul_comm] - -theorem inner_smul_ofReal_right (x y : F) {t : ℝ} : ⟪x, (t : 𝕜) • y⟫ = ⟪x, y⟫ * t := by - rw [inner_smul_right, mul_comm] - -theorem re_inner_smul_ofReal_smul_self (x : F) {t : ℝ} : - re ⟪(t : 𝕜) • x, (t : 𝕜) • x⟫ = normSqF x * t * t := by - apply ofReal_injective (K := 𝕜) - simp [inner_self_ofReal_re, inner_smul_ofReal_left, inner_smul_ofReal_right, normSq] - -/-- An auxiliary equality useful to prove the **Cauchy–Schwarz inequality**. Here we use the -standard argument involving the discriminant of quadratic form. -/ -lemma cauchy_schwarz_aux' (x y : F) (t : ℝ) : 0 ≤ normSqF x * t * t + 2 * re ⟪x, y⟫ * t - + normSqF y := by - calc 0 ≤ re ⟪(ofReal t : 𝕜) • x + y, (ofReal t : 𝕜) • x + y⟫ := inner_self_nonneg - _ = re (⟪(ofReal t : 𝕜) • x, (ofReal t : 𝕜) • x⟫ + ⟪(ofReal t : 𝕜) • x, y⟫ - + ⟪y, (ofReal t : 𝕜) • x⟫ + ⟪y, y⟫) := by rw [inner_add_add_self ((ofReal t : 𝕜) • x) y] - _ = re ⟪(ofReal t : 𝕜) • x, (ofReal t : 𝕜) • x⟫ - + re ⟪(ofReal t : 𝕜) • x, y⟫ + re ⟪y, (ofReal t : 𝕜) • x⟫ + re ⟪y, y⟫ := by - simp only [map_add] - _ = normSq x * t * t + re (⟪x, y⟫ * t) + re (⟪y, x⟫ * t) + re ⟪y, y⟫ := by rw - [re_inner_smul_ofReal_smul_self, inner_smul_ofReal_left, inner_smul_ofReal_right] - _ = normSq x * t * t + re ⟪x, y⟫ * t + re ⟪y, x⟫ * t + re ⟪y, y⟫ := by rw [mul_comm ⟪x,y⟫ _, - RCLike.re_ofReal_mul, mul_comm t _, mul_comm ⟪y,x⟫ _, RCLike.re_ofReal_mul, mul_comm t _] - _ = normSq x * t * t + re ⟪x, y⟫ * t + re ⟪y, x⟫ * t + normSq y := by rw [← normSq] - _ = normSq x * t * t + re ⟪x, y⟫ * t + re ⟪x, y⟫ * t + normSq y := by rw [inner_re_symm] - _ = normSq x * t * t + 2 * re ⟪x, y⟫ * t + normSq y := by ring - -/-- Another auxiliary equality related with the **Cauchy–Schwarz inequality**: the square of the -seminorm of `⟪x, y⟫ • x - ⟪x, x⟫ • y` is equal to `‖x‖ ^ 2 * (‖x‖ ^ 2 * ‖y‖ ^ 2 - ‖⟪x, y⟫‖ ^ 2)`. -We use `InnerProductSpace.ofCore.normSq x` etc (defeq to `is_R_or_C.re ⟪x, x⟫`) instead of `‖x‖ ^ 2` -etc to avoid extra rewrites when applying it to an `InnerProductSpace`. -/ -theorem cauchy_schwarz_aux (x y : F) : normSqF (⟪x, y⟫ • x - ⟪x, x⟫ • y) - = normSqF x * (normSqF x * normSqF y - ‖⟪x, y⟫‖ ^ 2) := by - rw [← @ofReal_inj 𝕜, ofReal_normSq_eq_inner_self] - simp only [inner_sub_sub_self, inner_smul_left, inner_smul_right, conj_ofReal, mul_sub, ← - ofReal_normSq_eq_inner_self x, ← ofReal_normSq_eq_inner_self y] - rw [← mul_assoc, mul_conj, RCLike.conj_mul, mul_left_comm, ← inner_conj_symm y, mul_conj] - push_cast - ring - -/-- **Cauchy–Schwarz inequality**. -We need this for the `PreInnerProductSpace.Core` structure to prove the triangle inequality below -when showing the core is a normed group and to take the quotient. --/ -theorem inner_mul_inner_self_le (x y : F) : ‖⟪x, y⟫‖ * ‖⟪y, x⟫‖ ≤ re ⟪x, x⟫ * re ⟪y, y⟫ := by - suffices discrim (normSqF x) (2 * ‖⟪x, y⟫_𝕜‖) (normSqF y) ≤ 0 by - rw [norm_inner_symm y x] - rw [discrim, normSq, normSq, sq] at this - linarith - refine discrim_le_zero fun t ↦ ?_ - by_cases hzero : ⟪x, y⟫ = 0 - · simp only [mul_assoc, ← sq, hzero, norm_zero, mul_zero, zero_mul, add_zero, ge_iff_le] - obtain ⟨hx, hy⟩ : (0 ≤ normSqF x ∧ 0 ≤ normSqF y) := ⟨inner_self_nonneg, inner_self_nonneg⟩ - positivity - · have hzero' : ‖⟪x, y⟫‖ ≠ 0 := norm_ne_zero_iff.2 hzero - convert cauchy_schwarz_aux' (𝕜 := 𝕜) (⟪x, y⟫ • x) y (t / ‖⟪x, y⟫‖) using 3 - · field_simp - rw [← sq, normSq, normSq, inner_smul_right, inner_smul_left, ← mul_assoc _ _ ⟪x, x⟫, - mul_conj] - nth_rw 2 [sq] - rw [← ofReal_mul, re_ofReal_mul] - ring - · field_simp - rw [inner_smul_left, mul_comm _ ⟪x, y⟫_𝕜, mul_conj, ← ofReal_pow, ofReal_re] - ring - -/-- (Semi)norm constructed from an `PreInnerProductSpace.Core` structure, defined to be the square -root of the scalar product. -/ -def toNorm : Norm F where norm x := √(re ⟪x, x⟫) - -attribute [local instance] toNorm - -theorem norm_eq_sqrt_inner (x : F) : ‖x‖ = √(re ⟪x, x⟫) := rfl - -theorem inner_self_eq_norm_mul_norm (x : F) : re ⟪x, x⟫ = ‖x‖ * ‖x‖ := by - rw [norm_eq_sqrt_inner, ← sqrt_mul inner_self_nonneg (re ⟪x, x⟫), sqrt_mul_self inner_self_nonneg] - -theorem sqrt_normSq_eq_norm (x : F) : √(normSqF x) = ‖x‖ := rfl - -/-- Cauchy–Schwarz inequality with norm -/ -theorem norm_inner_le_norm (x y : F) : ‖⟪x, y⟫‖ ≤ ‖x‖ * ‖y‖ := - nonneg_le_nonneg_of_sq_le_sq (mul_nonneg (sqrt_nonneg _) (sqrt_nonneg _)) <| - calc - ‖⟪x, y⟫‖ * ‖⟪x, y⟫‖ = ‖⟪x, y⟫‖ * ‖⟪y, x⟫‖ := by rw [norm_inner_symm] - _ ≤ re ⟪x, x⟫ * re ⟪y, y⟫ := inner_mul_inner_self_le x y - _ = ‖x‖ * ‖y‖ * (‖x‖ * ‖y‖) := by simp only [inner_self_eq_norm_mul_norm]; ring - -/-- Seminormed group structure constructed from an `PreInnerProductSpace.Core` structure -/ -def toSeminormedAddCommGroup : SeminormedAddCommGroup F := - AddGroupSeminorm.toSeminormedAddCommGroup - { toFun := fun x => √(re ⟪x, x⟫) - map_zero' := by simp only [sqrt_zero, inner_zero_right, map_zero] - neg' := fun x => by simp only [inner_neg_left, neg_neg, inner_neg_right] - add_le' := fun x y => by - have h₁ : ‖⟪x, y⟫‖ ≤ ‖x‖ * ‖y‖ := norm_inner_le_norm _ _ - have h₂ : re ⟪x, y⟫ ≤ ‖⟪x, y⟫‖ := re_le_norm _ - have h₃ : re ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := h₂.trans h₁ - have h₄ : re ⟪y, x⟫ ≤ ‖x‖ * ‖y‖ := by rwa [← inner_conj_symm, conj_re] - have : ‖x + y‖ * ‖x + y‖ ≤ (‖x‖ + ‖y‖) * (‖x‖ + ‖y‖) := by - simp only [← inner_self_eq_norm_mul_norm, inner_add_add_self, mul_add, mul_comm, map_add] - linarith - exact nonneg_le_nonneg_of_sq_le_sq (add_nonneg (sqrt_nonneg _) (sqrt_nonneg _)) this } - -attribute [local instance] toSeminormedAddCommGroup - -/-- Normed space (which is actually a seminorm) structure constructed from an -`PreInnerProductSpace.Core` structure -/ -def toSeminormedSpace : NormedSpace 𝕜 F where - norm_smul_le r x := by - rw [norm_eq_sqrt_inner, inner_smul_left, inner_smul_right, ← mul_assoc] - rw [RCLike.conj_mul, ← ofReal_pow, re_ofReal_mul, sqrt_mul, ← ofReal_normSq_eq_inner_self, - ofReal_re] - · simp [sqrt_normSq_eq_norm, RCLike.sqrt_normSq_eq_norm] - · positivity - -end PreInnerProductSpace.Core - -section InnerProductSpace.Core - -variable [AddCommGroup F] [Module 𝕜 F] [cd : InnerProductSpace.Core 𝕜 F] - -local notation "⟪" x ", " y "⟫" => @inner 𝕜 F _ x y - -local notation "normSqK" => @RCLike.normSq 𝕜 _ - -local notation "reK" => @RCLike.re 𝕜 _ - -local notation "ext_iff" => @RCLike.ext_iff 𝕜 _ - -local postfix:90 "†" => starRingEnd _ - -/-- Inner product defined by the `InnerProductSpace.Core` structure. We can't reuse -`InnerProductSpace.Core.toInner` because it takes `InnerProductSpace.Core` as an explicit -argument. -/ -def toInner' : Inner 𝕜 F := - cd.toInner - -attribute [local instance] toInner' - -local notation "normSqF" => @normSq 𝕜 F _ _ _ _ - -theorem inner_self_eq_zero {x : F} : ⟪x, x⟫ = 0 ↔ x = 0 := - ⟨cd.definite _, inner_self_of_eq_zero⟩ - -theorem normSq_eq_zero {x : F} : normSqF x = 0 ↔ x = 0 := - Iff.trans - (by simp only [normSq, ext_iff, map_zero, inner_self_im, eq_self_iff_true, and_true]) - (@inner_self_eq_zero 𝕜 _ _ _ _ _ x) - -theorem inner_self_ne_zero {x : F} : ⟪x, x⟫ ≠ 0 ↔ x ≠ 0 := - inner_self_eq_zero.not - -attribute [local instance] toNorm - -/-- Normed group structure constructed from an `InnerProductSpace.Core` structure -/ -def toNormedAddCommGroup : NormedAddCommGroup F := - AddGroupNorm.toNormedAddCommGroup - { toFun := fun x => √(re ⟪x, x⟫) - map_zero' := by simp only [sqrt_zero, inner_zero_right, map_zero] - neg' := fun x => by simp only [inner_neg_left, neg_neg, inner_neg_right] - add_le' := fun x y => by - have h₁ : ‖⟪x, y⟫‖ ≤ ‖x‖ * ‖y‖ := norm_inner_le_norm _ _ - have h₂ : re ⟪x, y⟫ ≤ ‖⟪x, y⟫‖ := re_le_norm _ - have h₃ : re ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := h₂.trans h₁ - have h₄ : re ⟪y, x⟫ ≤ ‖x‖ * ‖y‖ := by rwa [← inner_conj_symm, conj_re] - have : ‖x + y‖ * ‖x + y‖ ≤ (‖x‖ + ‖y‖) * (‖x‖ + ‖y‖) := by - simp only [← inner_self_eq_norm_mul_norm, inner_add_add_self, mul_add, mul_comm, map_add] - linarith - exact nonneg_le_nonneg_of_sq_le_sq (add_nonneg (sqrt_nonneg _) (sqrt_nonneg _)) this - eq_zero_of_map_eq_zero' := fun _ hx => - normSq_eq_zero.1 <| (sqrt_eq_zero inner_self_nonneg).1 hx } - -attribute [local instance] toNormedAddCommGroup - -/-- Normed space structure constructed from an `InnerProductSpace.Core` structure -/ -def toNormedSpace : NormedSpace 𝕜 F where - norm_smul_le r x := by - rw [norm_eq_sqrt_inner, inner_smul_left, inner_smul_right, ← mul_assoc] - rw [RCLike.conj_mul, ← ofReal_pow, re_ofReal_mul, sqrt_mul, ← ofReal_normSq_eq_inner_self, - ofReal_re] - · simp [sqrt_normSq_eq_norm, RCLike.sqrt_normSq_eq_norm] - · positivity - -end InnerProductSpace.Core - -end InnerProductSpace.Core - -section - -attribute [local instance] InnerProductSpace.Core.toNormedAddCommGroup - -/-- Given an `InnerProductSpace.Core` structure on a space, one can use it to turn -the space into an inner product space. The `NormedAddCommGroup` structure is expected -to already be defined with `InnerProductSpace.ofCore.toNormedAddCommGroup`. -/ -def InnerProductSpace.ofCore [AddCommGroup F] [Module 𝕜 F] (cd : InnerProductSpace.Core 𝕜 F) : - InnerProductSpace 𝕜 F := - letI : NormedSpace 𝕜 F := @InnerProductSpace.Core.toNormedSpace 𝕜 F _ _ _ cd - { cd with - norm_sq_eq_inner := fun x => by - have h₁ : ‖x‖ ^ 2 = √(re (cd.inner x x)) ^ 2 := rfl - have h₂ : 0 ≤ re (cd.inner x x) := InnerProductSpace.Core.inner_self_nonneg - simp [h₁, sq_sqrt, h₂] } - -end - -/-! ### Properties of inner product spaces -/ - section BasicProperties_Seminormed open scoped InnerProductSpace @@ -517,8 +47,6 @@ variable [SeminormedAddCommGroup F] [InnerProductSpace ℝ F] local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y -local notation "IK" => @RCLike.I 𝕜 _ - local postfix:90 "†" => starRingEnd _ export InnerProductSpace (norm_sq_eq_inner) @@ -741,7 +269,6 @@ theorem real_inner_mul_inner_self_le (x y : F) : ⟪x, y⟫_ℝ * ⟪x, y⟫_ℝ exact le_abs_self _ _ ≤ ⟪x, x⟫_ℝ * ⟪y, y⟫_ℝ := @inner_mul_inner_self_le ℝ _ _ _ _ x y - end BasicProperties_Seminormed section BasicProperties @@ -751,10 +278,6 @@ variable [NormedAddCommGroup F] [InnerProductSpace ℝ F] local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y -local notation "IK" => @RCLike.I 𝕜 _ - -local postfix:90 "†" => starRingEnd _ - export InnerProductSpace (norm_sq_eq_inner) @[simp] @@ -800,239 +323,6 @@ theorem linearIndependent_of_ne_zero_of_inner_eq_zero {ι : Type*} {v : ι → E end BasicProperties -section OrthonormalSets_Seminormed - -variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] -variable [SeminormedAddCommGroup F] [InnerProductSpace ℝ F] - -local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y - -local notation "IK" => @RCLike.I 𝕜 _ - -local postfix:90 "†" => starRingEnd _ - -variable {ι : Type*} (𝕜) - -/-- An orthonormal set of vectors in an `InnerProductSpace` -/ -def Orthonormal (v : ι → E) : Prop := - (∀ i, ‖v i‖ = 1) ∧ Pairwise fun i j => ⟪v i, v j⟫ = 0 - -variable {𝕜} - -/-- `if ... then ... else` characterization of an indexed set of vectors being orthonormal. (Inner -product equals Kronecker delta.) -/ -theorem orthonormal_iff_ite [DecidableEq ι] {v : ι → E} : - Orthonormal 𝕜 v ↔ ∀ i j, ⟪v i, v j⟫ = if i = j then (1 : 𝕜) else (0 : 𝕜) := by - constructor - · intro hv i j - split_ifs with h - · simp [h, inner_self_eq_norm_sq_to_K, hv.1] - · exact hv.2 h - · intro h - constructor - · intro i - have h' : ‖v i‖ ^ 2 = 1 ^ 2 := by simp [@norm_sq_eq_inner 𝕜, h i i] - have h₁ : 0 ≤ ‖v i‖ := norm_nonneg _ - have h₂ : (0 : ℝ) ≤ 1 := zero_le_one - rwa [sq_eq_sq₀ h₁ h₂] at h' - · intro i j hij - simpa [hij] using h i j - -/-- `if ... then ... else` characterization of a set of vectors being orthonormal. (Inner product -equals Kronecker delta.) -/ -theorem orthonormal_subtype_iff_ite [DecidableEq E] {s : Set E} : - Orthonormal 𝕜 (Subtype.val : s → E) ↔ ∀ v ∈ s, ∀ w ∈ s, ⟪v, w⟫ = if v = w then 1 else 0 := by - rw [orthonormal_iff_ite] - constructor - · intro h v hv w hw - convert h ⟨v, hv⟩ ⟨w, hw⟩ using 1 - simp - · rintro h ⟨v, hv⟩ ⟨w, hw⟩ - convert h v hv w hw using 1 - simp - -/-- The inner product of a linear combination of a set of orthonormal vectors with one of those -vectors picks out the coefficient of that vector. -/ -theorem Orthonormal.inner_right_finsupp {v : ι → E} (hv : Orthonormal 𝕜 v) (l : ι →₀ 𝕜) (i : ι) : - ⟪v i, linearCombination 𝕜 v l⟫ = l i := by - classical - simpa [linearCombination_apply, Finsupp.inner_sum, orthonormal_iff_ite.mp hv] using Eq.symm - -/-- The inner product of a linear combination of a set of orthonormal vectors with one of those -vectors picks out the coefficient of that vector. -/ -theorem Orthonormal.inner_right_sum {v : ι → E} (hv : Orthonormal 𝕜 v) (l : ι → 𝕜) {s : Finset ι} - {i : ι} (hi : i ∈ s) : ⟪v i, ∑ i ∈ s, l i • v i⟫ = l i := by - classical - simp [inner_sum, inner_smul_right, orthonormal_iff_ite.mp hv, hi] - -/-- The inner product of a linear combination of a set of orthonormal vectors with one of those -vectors picks out the coefficient of that vector. -/ -theorem Orthonormal.inner_right_fintype [Fintype ι] {v : ι → E} (hv : Orthonormal 𝕜 v) (l : ι → 𝕜) - (i : ι) : ⟪v i, ∑ i : ι, l i • v i⟫ = l i := - hv.inner_right_sum l (Finset.mem_univ _) - -/-- The inner product of a linear combination of a set of orthonormal vectors with one of those -vectors picks out the coefficient of that vector. -/ -theorem Orthonormal.inner_left_finsupp {v : ι → E} (hv : Orthonormal 𝕜 v) (l : ι →₀ 𝕜) (i : ι) : - ⟪linearCombination 𝕜 v l, v i⟫ = conj (l i) := by rw [← inner_conj_symm, hv.inner_right_finsupp] - -/-- The inner product of a linear combination of a set of orthonormal vectors with one of those -vectors picks out the coefficient of that vector. -/ -theorem Orthonormal.inner_left_sum {v : ι → E} (hv : Orthonormal 𝕜 v) (l : ι → 𝕜) {s : Finset ι} - {i : ι} (hi : i ∈ s) : ⟪∑ i ∈ s, l i • v i, v i⟫ = conj (l i) := by - classical - simp only [sum_inner, inner_smul_left, orthonormal_iff_ite.mp hv, hi, mul_boole, - Finset.sum_ite_eq', if_true] - -/-- The inner product of a linear combination of a set of orthonormal vectors with one of those -vectors picks out the coefficient of that vector. -/ -theorem Orthonormal.inner_left_fintype [Fintype ι] {v : ι → E} (hv : Orthonormal 𝕜 v) (l : ι → 𝕜) - (i : ι) : ⟪∑ i : ι, l i • v i, v i⟫ = conj (l i) := - hv.inner_left_sum l (Finset.mem_univ _) - -/-- The inner product of two linear combinations of a set of orthonormal vectors, expressed as -a sum over the first `Finsupp`. -/ -theorem Orthonormal.inner_finsupp_eq_sum_left {v : ι → E} (hv : Orthonormal 𝕜 v) (l₁ l₂ : ι →₀ 𝕜) : - ⟪linearCombination 𝕜 v l₁, linearCombination 𝕜 v l₂⟫ = l₁.sum fun i y => conj y * l₂ i := by - simp only [l₁.linearCombination_apply _, Finsupp.sum_inner, hv.inner_right_finsupp, smul_eq_mul] - -/-- The inner product of two linear combinations of a set of orthonormal vectors, expressed as -a sum over the second `Finsupp`. -/ -theorem Orthonormal.inner_finsupp_eq_sum_right {v : ι → E} (hv : Orthonormal 𝕜 v) (l₁ l₂ : ι →₀ 𝕜) : - ⟪linearCombination 𝕜 v l₁, linearCombination 𝕜 v l₂⟫ = l₂.sum fun i y => conj (l₁ i) * y := by - simp only [l₂.linearCombination_apply _, Finsupp.inner_sum, hv.inner_left_finsupp, mul_comm, - smul_eq_mul] - -/-- The inner product of two linear combinations of a set of orthonormal vectors, expressed as -a sum. -/ -protected theorem Orthonormal.inner_sum {v : ι → E} (hv : Orthonormal 𝕜 v) (l₁ l₂ : ι → 𝕜) - (s : Finset ι) : ⟪∑ i ∈ s, l₁ i • v i, ∑ i ∈ s, l₂ i • v i⟫ = ∑ i ∈ s, conj (l₁ i) * l₂ i := by - simp_rw [sum_inner, inner_smul_left] - refine Finset.sum_congr rfl fun i hi => ?_ - rw [hv.inner_right_sum l₂ hi] - -/-- -The double sum of weighted inner products of pairs of vectors from an orthonormal sequence is the -sum of the weights. --/ -theorem Orthonormal.inner_left_right_finset {s : Finset ι} {v : ι → E} (hv : Orthonormal 𝕜 v) - {a : ι → ι → 𝕜} : (∑ i ∈ s, ∑ j ∈ s, a i j • ⟪v j, v i⟫) = ∑ k ∈ s, a k k := by - classical - simp [orthonormal_iff_ite.mp hv, Finset.sum_ite_of_true] - -/-- An orthonormal set is linearly independent. -/ -theorem Orthonormal.linearIndependent {v : ι → E} (hv : Orthonormal 𝕜 v) : - LinearIndependent 𝕜 v := by - rw [linearIndependent_iff] - intro l hl - ext i - have key : ⟪v i, Finsupp.linearCombination 𝕜 v l⟫ = ⟪v i, 0⟫ := by rw [hl] - simpa only [hv.inner_right_finsupp, inner_zero_right] using key - -/-- A subfamily of an orthonormal family (i.e., a composition with an injective map) is an -orthonormal family. -/ -theorem Orthonormal.comp {ι' : Type*} {v : ι → E} (hv : Orthonormal 𝕜 v) (f : ι' → ι) - (hf : Function.Injective f) : Orthonormal 𝕜 (v ∘ f) := by - classical - rw [orthonormal_iff_ite] at hv ⊢ - intro i j - convert hv (f i) (f j) using 1 - simp [hf.eq_iff] - -/-- An injective family `v : ι → E` is orthonormal if and only if `Subtype.val : (range v) → E` is -orthonormal. -/ -theorem orthonormal_subtype_range {v : ι → E} (hv : Function.Injective v) : - Orthonormal 𝕜 (Subtype.val : Set.range v → E) ↔ Orthonormal 𝕜 v := by - let f : ι ≃ Set.range v := Equiv.ofInjective v hv - refine ⟨fun h => h.comp f f.injective, fun h => ?_⟩ - rw [← Equiv.self_comp_ofInjective_symm hv] - exact h.comp f.symm f.symm.injective - -/-- If `v : ι → E` is an orthonormal family, then `Subtype.val : (range v) → E` is an orthonormal -family. -/ -theorem Orthonormal.toSubtypeRange {v : ι → E} (hv : Orthonormal 𝕜 v) : - Orthonormal 𝕜 (Subtype.val : Set.range v → E) := - (orthonormal_subtype_range hv.linearIndependent.injective).2 hv - -/-- A linear combination of some subset of an orthonormal set is orthogonal to other members of the -set. -/ -theorem Orthonormal.inner_finsupp_eq_zero {v : ι → E} (hv : Orthonormal 𝕜 v) {s : Set ι} {i : ι} - (hi : i ∉ s) {l : ι →₀ 𝕜} (hl : l ∈ Finsupp.supported 𝕜 𝕜 s) : - ⟪Finsupp.linearCombination 𝕜 v l, v i⟫ = 0 := by - rw [Finsupp.mem_supported'] at hl - simp only [hv.inner_left_finsupp, hl i hi, map_zero] - -/-- Given an orthonormal family, a second family of vectors is orthonormal if every vector equals -the corresponding vector in the original family or its negation. -/ -theorem Orthonormal.orthonormal_of_forall_eq_or_eq_neg {v w : ι → E} (hv : Orthonormal 𝕜 v) - (hw : ∀ i, w i = v i ∨ w i = -v i) : Orthonormal 𝕜 w := by - classical - rw [orthonormal_iff_ite] at * - intro i j - cases' hw i with hi hi <;> cases' hw j with hj hj <;> - replace hv := hv i j <;> split_ifs at hv ⊢ with h <;> - simpa only [hi, hj, h, inner_neg_right, inner_neg_left, neg_neg, eq_self_iff_true, - neg_eq_zero] using hv - -/- The material that follows, culminating in the existence of a maximal orthonormal subset, is -adapted from the corresponding development of the theory of linearly independents sets. See -`exists_linearIndependent` in particular. -/ -variable (𝕜 E) - -theorem orthonormal_empty : Orthonormal 𝕜 (fun x => x : (∅ : Set E) → E) := by - classical - simp [orthonormal_subtype_iff_ite] - -variable {𝕜 E} - -theorem orthonormal_iUnion_of_directed {η : Type*} {s : η → Set E} (hs : Directed (· ⊆ ·) s) - (h : ∀ i, Orthonormal 𝕜 (fun x => x : s i → E)) : - Orthonormal 𝕜 (fun x => x : (⋃ i, s i) → E) := by - classical - rw [orthonormal_subtype_iff_ite] - rintro x ⟨_, ⟨i, rfl⟩, hxi⟩ y ⟨_, ⟨j, rfl⟩, hyj⟩ - obtain ⟨k, hik, hjk⟩ := hs i j - have h_orth : Orthonormal 𝕜 (fun x => x : s k → E) := h k - rw [orthonormal_subtype_iff_ite] at h_orth - exact h_orth x (hik hxi) y (hjk hyj) - -theorem orthonormal_sUnion_of_directed {s : Set (Set E)} (hs : DirectedOn (· ⊆ ·) s) - (h : ∀ a ∈ s, Orthonormal 𝕜 (fun x => ((x : a) : E))) : - Orthonormal 𝕜 (fun x => x : ⋃₀ s → E) := by - rw [Set.sUnion_eq_iUnion]; exact orthonormal_iUnion_of_directed hs.directed_val (by simpa using h) - -/-- Given an orthonormal set `v` of vectors in `E`, there exists a maximal orthonormal set -containing it. -/ -theorem exists_maximal_orthonormal {s : Set E} (hs : Orthonormal 𝕜 (Subtype.val : s → E)) : - ∃ w ⊇ s, Orthonormal 𝕜 (Subtype.val : w → E) ∧ - ∀ u ⊇ w, Orthonormal 𝕜 (Subtype.val : u → E) → u = w := by - have := zorn_subset_nonempty { b | Orthonormal 𝕜 (Subtype.val : b → E) } ?_ _ hs - · obtain ⟨b, hb⟩ := this - exact ⟨b, hb.1, hb.2.1, fun u hus hu => hb.2.eq_of_ge hu hus ⟩ - · refine fun c hc cc _c0 => ⟨⋃₀ c, ?_, ?_⟩ - · exact orthonormal_sUnion_of_directed cc.directedOn fun x xc => hc xc - · exact fun _ => Set.subset_sUnion_of_mem - -open Module - -/-- A family of orthonormal vectors with the correct cardinality forms a basis. -/ -def basisOfOrthonormalOfCardEqFinrank [Fintype ι] [Nonempty ι] {v : ι → E} (hv : Orthonormal 𝕜 v) - (card_eq : Fintype.card ι = finrank 𝕜 E) : Basis ι 𝕜 E := - basisOfLinearIndependentOfCardEqFinrank hv.linearIndependent card_eq - -@[simp] -theorem coe_basisOfOrthonormalOfCardEqFinrank [Fintype ι] [Nonempty ι] {v : ι → E} - (hv : Orthonormal 𝕜 v) (card_eq : Fintype.card ι = finrank 𝕜 E) : - (basisOfOrthonormalOfCardEqFinrank hv card_eq : ι → E) = v := - coe_basisOfLinearIndependentOfCardEqFinrank _ _ - -theorem Orthonormal.ne_zero {v : ι → E} (hv : Orthonormal 𝕜 v) (i : ι) : v i ≠ 0 := by - refine ne_of_apply_ne norm ?_ - rw [hv.1 i, norm_zero] - norm_num - -end OrthonormalSets_Seminormed - section Norm_Seminormed open scoped InnerProductSpace @@ -1044,8 +334,6 @@ local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y local notation "IK" => @RCLike.I 𝕜 _ -local postfix:90 "†" => starRingEnd _ - theorem norm_eq_sqrt_inner (x : E) : ‖x‖ = √(re ⟪x, x⟫) := calc ‖x‖ = √(‖x‖ ^ 2) := (sqrt_sq (norm_nonneg _)).symm @@ -1214,214 +502,6 @@ instance (priority := 100) InnerProductSpace.toUniformConvexSpace : UniformConve ring_nf gcongr⟩ -section Complex_Seminormed - -variable {V : Type*} [SeminormedAddCommGroup V] [InnerProductSpace ℂ V] - -/-- A complex polarization identity, with a linear map --/ -theorem inner_map_polarization (T : V →ₗ[ℂ] V) (x y : V) : - ⟪T y, x⟫_ℂ = - (⟪T (x + y), x + y⟫_ℂ - ⟪T (x - y), x - y⟫_ℂ + - Complex.I * ⟪T (x + Complex.I • y), x + Complex.I • y⟫_ℂ - - Complex.I * ⟪T (x - Complex.I • y), x - Complex.I • y⟫_ℂ) / - 4 := by - simp only [map_add, map_sub, inner_add_left, inner_add_right, LinearMap.map_smul, inner_smul_left, - inner_smul_right, Complex.conj_I, ← pow_two, Complex.I_sq, inner_sub_left, inner_sub_right, - mul_add, ← mul_assoc, mul_neg, neg_neg, sub_neg_eq_add, one_mul, neg_one_mul, mul_sub, sub_sub] - ring - -theorem inner_map_polarization' (T : V →ₗ[ℂ] V) (x y : V) : - ⟪T x, y⟫_ℂ = - (⟪T (x + y), x + y⟫_ℂ - ⟪T (x - y), x - y⟫_ℂ - - Complex.I * ⟪T (x + Complex.I • y), x + Complex.I • y⟫_ℂ + - Complex.I * ⟪T (x - Complex.I • y), x - Complex.I • y⟫_ℂ) / - 4 := by - simp only [map_add, map_sub, inner_add_left, inner_add_right, LinearMap.map_smul, inner_smul_left, - inner_smul_right, Complex.conj_I, ← pow_two, Complex.I_sq, inner_sub_left, inner_sub_right, - mul_add, ← mul_assoc, mul_neg, neg_neg, sub_neg_eq_add, one_mul, neg_one_mul, mul_sub, sub_sub] - ring - -end Complex_Seminormed - -section Complex - -variable {V : Type*} [NormedAddCommGroup V] [InnerProductSpace ℂ V] - -/-- A linear map `T` is zero, if and only if the identity `⟪T x, x⟫_ℂ = 0` holds for all `x`. --/ -theorem inner_map_self_eq_zero (T : V →ₗ[ℂ] V) : (∀ x : V, ⟪T x, x⟫_ℂ = 0) ↔ T = 0 := by - constructor - · intro hT - ext x - rw [LinearMap.zero_apply, ← @inner_self_eq_zero ℂ V, inner_map_polarization] - simp only [hT] - norm_num - · rintro rfl x - simp only [LinearMap.zero_apply, inner_zero_left] - -/-- -Two linear maps `S` and `T` are equal, if and only if the identity `⟪S x, x⟫_ℂ = ⟪T x, x⟫_ℂ` holds -for all `x`. --/ -theorem ext_inner_map (S T : V →ₗ[ℂ] V) : (∀ x : V, ⟪S x, x⟫_ℂ = ⟪T x, x⟫_ℂ) ↔ S = T := by - rw [← sub_eq_zero, ← inner_map_self_eq_zero] - refine forall_congr' fun x => ?_ - rw [LinearMap.sub_apply, inner_sub_left, sub_eq_zero] - -end Complex - -section - -variable {ι : Type*} {ι' : Type*} {ι'' : Type*} -variable {E' : Type*} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] -variable {E'' : Type*} [SeminormedAddCommGroup E''] [InnerProductSpace 𝕜 E''] - -/-- A linear isometry preserves the inner product. -/ -@[simp] -theorem LinearIsometry.inner_map_map (f : E →ₗᵢ[𝕜] E') (x y : E) : ⟪f x, f y⟫ = ⟪x, y⟫ := by - simp [inner_eq_sum_norm_sq_div_four, ← f.norm_map] - -/-- A linear isometric equivalence preserves the inner product. -/ -@[simp] -theorem LinearIsometryEquiv.inner_map_map (f : E ≃ₗᵢ[𝕜] E') (x y : E) : ⟪f x, f y⟫ = ⟪x, y⟫ := - f.toLinearIsometry.inner_map_map x y - -/-- The adjoint of a linear isometric equivalence is its inverse. -/ -theorem LinearIsometryEquiv.inner_map_eq_flip (f : E ≃ₗᵢ[𝕜] E') (x : E) (y : E') : - ⟪f x, y⟫_𝕜 = ⟪x, f.symm y⟫_𝕜 := by - conv_lhs => rw [← f.apply_symm_apply y, f.inner_map_map] - -/-- A linear map that preserves the inner product is a linear isometry. -/ -def LinearMap.isometryOfInner (f : E →ₗ[𝕜] E') (h : ∀ x y, ⟪f x, f y⟫ = ⟪x, y⟫) : E →ₗᵢ[𝕜] E' := - ⟨f, fun x => by simp only [@norm_eq_sqrt_inner 𝕜, h]⟩ - -@[simp] -theorem LinearMap.coe_isometryOfInner (f : E →ₗ[𝕜] E') (h) : ⇑(f.isometryOfInner h) = f := - rfl - -@[simp] -theorem LinearMap.isometryOfInner_toLinearMap (f : E →ₗ[𝕜] E') (h) : - (f.isometryOfInner h).toLinearMap = f := - rfl - -/-- A linear equivalence that preserves the inner product is a linear isometric equivalence. -/ -def LinearEquiv.isometryOfInner (f : E ≃ₗ[𝕜] E') (h : ∀ x y, ⟪f x, f y⟫ = ⟪x, y⟫) : E ≃ₗᵢ[𝕜] E' := - ⟨f, ((f : E →ₗ[𝕜] E').isometryOfInner h).norm_map⟩ - -@[simp] -theorem LinearEquiv.coe_isometryOfInner (f : E ≃ₗ[𝕜] E') (h) : ⇑(f.isometryOfInner h) = f := - rfl - -@[simp] -theorem LinearEquiv.isometryOfInner_toLinearEquiv (f : E ≃ₗ[𝕜] E') (h) : - (f.isometryOfInner h).toLinearEquiv = f := - rfl - -/-- A linear map is an isometry if and it preserves the inner product. -/ -theorem LinearMap.norm_map_iff_inner_map_map {F : Type*} [FunLike F E E'] [LinearMapClass F 𝕜 E E'] - (f : F) : (∀ x, ‖f x‖ = ‖x‖) ↔ (∀ x y, ⟪f x, f y⟫_𝕜 = ⟪x, y⟫_𝕜) := - ⟨({ toLinearMap := LinearMapClass.linearMap f, norm_map' := · : E →ₗᵢ[𝕜] E' }.inner_map_map), - (LinearMapClass.linearMap f |>.isometryOfInner · |>.norm_map)⟩ - -/-- A linear isometry preserves the property of being orthonormal. -/ -theorem LinearIsometry.orthonormal_comp_iff {v : ι → E} (f : E →ₗᵢ[𝕜] E') : - Orthonormal 𝕜 (f ∘ v) ↔ Orthonormal 𝕜 v := by - classical simp_rw [orthonormal_iff_ite, Function.comp_apply, LinearIsometry.inner_map_map] - -/-- A linear isometry preserves the property of being orthonormal. -/ -theorem Orthonormal.comp_linearIsometry {v : ι → E} (hv : Orthonormal 𝕜 v) (f : E →ₗᵢ[𝕜] E') : - Orthonormal 𝕜 (f ∘ v) := by rwa [f.orthonormal_comp_iff] - -/-- A linear isometric equivalence preserves the property of being orthonormal. -/ -theorem Orthonormal.comp_linearIsometryEquiv {v : ι → E} (hv : Orthonormal 𝕜 v) (f : E ≃ₗᵢ[𝕜] E') : - Orthonormal 𝕜 (f ∘ v) := - hv.comp_linearIsometry f.toLinearIsometry - -/-- A linear isometric equivalence, applied with `Basis.map`, preserves the property of being -orthonormal. -/ -theorem Orthonormal.mapLinearIsometryEquiv {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) - (f : E ≃ₗᵢ[𝕜] E') : Orthonormal 𝕜 (v.map f.toLinearEquiv) := - hv.comp_linearIsometryEquiv f - -/-- A linear map that sends an orthonormal basis to orthonormal vectors is a linear isometry. -/ -def LinearMap.isometryOfOrthonormal (f : E →ₗ[𝕜] E') {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) - (hf : Orthonormal 𝕜 (f ∘ v)) : E →ₗᵢ[𝕜] E' := - f.isometryOfInner fun x y => by - classical rw [← v.linearCombination_repr x, ← v.linearCombination_repr y, - Finsupp.apply_linearCombination, Finsupp.apply_linearCombination, - hv.inner_finsupp_eq_sum_left, hf.inner_finsupp_eq_sum_left] - -@[simp] -theorem LinearMap.coe_isometryOfOrthonormal (f : E →ₗ[𝕜] E') {v : Basis ι 𝕜 E} - (hv : Orthonormal 𝕜 v) (hf : Orthonormal 𝕜 (f ∘ v)) : ⇑(f.isometryOfOrthonormal hv hf) = f := - rfl - -@[simp] -theorem LinearMap.isometryOfOrthonormal_toLinearMap (f : E →ₗ[𝕜] E') {v : Basis ι 𝕜 E} - (hv : Orthonormal 𝕜 v) (hf : Orthonormal 𝕜 (f ∘ v)) : - (f.isometryOfOrthonormal hv hf).toLinearMap = f := - rfl - -/-- A linear equivalence that sends an orthonormal basis to orthonormal vectors is a linear -isometric equivalence. -/ -def LinearEquiv.isometryOfOrthonormal (f : E ≃ₗ[𝕜] E') {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) - (hf : Orthonormal 𝕜 (f ∘ v)) : E ≃ₗᵢ[𝕜] E' := - f.isometryOfInner fun x y => by - rw [← LinearEquiv.coe_coe] at hf - classical rw [← v.linearCombination_repr x, ← v.linearCombination_repr y, - ← LinearEquiv.coe_coe f, Finsupp.apply_linearCombination, - Finsupp.apply_linearCombination, hv.inner_finsupp_eq_sum_left, hf.inner_finsupp_eq_sum_left] - -@[simp] -theorem LinearEquiv.coe_isometryOfOrthonormal (f : E ≃ₗ[𝕜] E') {v : Basis ι 𝕜 E} - (hv : Orthonormal 𝕜 v) (hf : Orthonormal 𝕜 (f ∘ v)) : ⇑(f.isometryOfOrthonormal hv hf) = f := - rfl - -@[simp] -theorem LinearEquiv.isometryOfOrthonormal_toLinearEquiv (f : E ≃ₗ[𝕜] E') {v : Basis ι 𝕜 E} - (hv : Orthonormal 𝕜 v) (hf : Orthonormal 𝕜 (f ∘ v)) : - (f.isometryOfOrthonormal hv hf).toLinearEquiv = f := - rfl - -/-- A linear isometric equivalence that sends an orthonormal basis to a given orthonormal basis. -/ -def Orthonormal.equiv {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) {v' : Basis ι' 𝕜 E'} - (hv' : Orthonormal 𝕜 v') (e : ι ≃ ι') : E ≃ₗᵢ[𝕜] E' := - (v.equiv v' e).isometryOfOrthonormal hv - (by - have h : v.equiv v' e ∘ v = v' ∘ e := by - ext i - simp - rw [h] - classical exact hv'.comp _ e.injective) - -@[simp] -theorem Orthonormal.equiv_toLinearEquiv {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) - {v' : Basis ι' 𝕜 E'} (hv' : Orthonormal 𝕜 v') (e : ι ≃ ι') : - (hv.equiv hv' e).toLinearEquiv = v.equiv v' e := - rfl - -@[simp] -theorem Orthonormal.equiv_apply {ι' : Type*} {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) - {v' : Basis ι' 𝕜 E'} (hv' : Orthonormal 𝕜 v') (e : ι ≃ ι') (i : ι) : - hv.equiv hv' e (v i) = v' (e i) := - Basis.equiv_apply _ _ _ _ - -@[simp] -theorem Orthonormal.equiv_trans {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) {v' : Basis ι' 𝕜 E'} - (hv' : Orthonormal 𝕜 v') (e : ι ≃ ι') {v'' : Basis ι'' 𝕜 E''} (hv'' : Orthonormal 𝕜 v'') - (e' : ι' ≃ ι'') : (hv.equiv hv' e).trans (hv'.equiv hv'' e') = hv.equiv hv'' (e.trans e') := - v.ext_linearIsometryEquiv fun i => by - simp only [LinearIsometryEquiv.trans_apply, Orthonormal.equiv_apply, e.coe_trans, - Function.comp_apply] - -theorem Orthonormal.map_equiv {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) {v' : Basis ι' 𝕜 E'} - (hv' : Orthonormal 𝕜 v') (e : ι ≃ ι') : - v.map (hv.equiv hv' e).toLinearEquiv = v'.reindex e.symm := - v.map_equiv _ _ - -end - /-- Polarization identity: The real inner product, in terms of the norm. -/ theorem real_inner_eq_norm_add_mul_self_sub_norm_mul_self_sub_norm_mul_self_div_two (x y : F) : ⟪x, y⟫_ℝ = (‖x + y‖ * ‖x + y‖ - ‖x‖ * ‖x‖ - ‖y‖ * ‖y‖) / 2 := @@ -1513,131 +593,6 @@ theorem real_inner_smul_self_left (x : F) (r : ℝ) : ⟪r • x, x⟫_ℝ = r * theorem real_inner_smul_self_right (x : F) (r : ℝ) : ⟪x, r • x⟫_ℝ = r * (‖x‖ * ‖x‖) := by rw [inner_smul_right, ← real_inner_self_eq_norm_mul_norm] -variable (𝕜) - -/-- The inner product as a sesquilinear map. -/ -def innerₛₗ : E →ₗ⋆[𝕜] E →ₗ[𝕜] 𝕜 := - LinearMap.mk₂'ₛₗ _ _ (fun v w => ⟪v, w⟫) inner_add_left (fun _ _ _ => inner_smul_left _ _ _) - inner_add_right fun _ _ _ => inner_smul_right _ _ _ - -@[simp] -theorem innerₛₗ_apply_coe (v : E) : ⇑(innerₛₗ 𝕜 v) = fun w => ⟪v, w⟫ := - rfl - -@[simp] -theorem innerₛₗ_apply (v w : E) : innerₛₗ 𝕜 v w = ⟪v, w⟫ := - rfl - -variable (F) -/-- The inner product as a bilinear map in the real case. -/ -def innerₗ : F →ₗ[ℝ] F →ₗ[ℝ] ℝ := innerₛₗ ℝ - -@[simp] lemma flip_innerₗ : (innerₗ F).flip = innerₗ F := by - ext v w - exact real_inner_comm v w - -variable {F} - -@[simp] lemma innerₗ_apply (v w : F) : innerₗ F v w = ⟪v, w⟫_ℝ := rfl - -/-- The inner product as a continuous sesquilinear map. Note that `toDualMap` (resp. `toDual`) -in `InnerProductSpace.Dual` is a version of this given as a linear isometry (resp. linear -isometric equivalence). -/ -def innerSL : E →L⋆[𝕜] E →L[𝕜] 𝕜 := - LinearMap.mkContinuous₂ (innerₛₗ 𝕜) 1 fun x y => by - simp only [norm_inner_le_norm, one_mul, innerₛₗ_apply] - -@[simp] -theorem innerSL_apply_coe (v : E) : ⇑(innerSL 𝕜 v) = fun w => ⟪v, w⟫ := - rfl - -@[simp] -theorem innerSL_apply (v w : E) : innerSL 𝕜 v w = ⟪v, w⟫ := - rfl - -/-- The inner product as a continuous sesquilinear map, with the two arguments flipped. -/ -def innerSLFlip : E →L[𝕜] E →L⋆[𝕜] 𝕜 := - @ContinuousLinearMap.flipₗᵢ' 𝕜 𝕜 𝕜 E E 𝕜 _ _ _ _ _ _ _ _ _ (RingHom.id 𝕜) (starRingEnd 𝕜) _ _ - (innerSL 𝕜) - -@[simp] -theorem innerSLFlip_apply (x y : E) : innerSLFlip 𝕜 x y = ⟪y, x⟫ := - rfl - -variable (F) in -@[simp] lemma innerSL_real_flip : (innerSL ℝ (E := F)).flip = innerSL ℝ := by - ext v w - exact real_inner_comm _ _ - -variable {𝕜} - -namespace ContinuousLinearMap - -variable {E' : Type*} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] - --- Note: odd and expensive build behavior is explicitly turned off using `noncomputable` -/-- Given `f : E →L[𝕜] E'`, construct the continuous sesquilinear form `fun x y ↦ ⟪x, A y⟫`, given -as a continuous linear map. -/ -noncomputable def toSesqForm : (E →L[𝕜] E') →L[𝕜] E' →L⋆[𝕜] E →L[𝕜] 𝕜 := - (ContinuousLinearMap.flipₗᵢ' E E' 𝕜 (starRingEnd 𝕜) (RingHom.id 𝕜)).toContinuousLinearEquiv ∘L - ContinuousLinearMap.compSL E E' (E' →L⋆[𝕜] 𝕜) (RingHom.id 𝕜) (RingHom.id 𝕜) (innerSLFlip 𝕜) - -@[simp] -theorem toSesqForm_apply_coe (f : E →L[𝕜] E') (x : E') : toSesqForm f x = (innerSL 𝕜 x).comp f := - rfl - -theorem toSesqForm_apply_norm_le {f : E →L[𝕜] E'} {v : E'} : ‖toSesqForm f v‖ ≤ ‖f‖ * ‖v‖ := by - refine opNorm_le_bound _ (by positivity) fun x ↦ ?_ - have h₁ : ‖f x‖ ≤ ‖f‖ * ‖x‖ := le_opNorm _ _ - have h₂ := @norm_inner_le_norm 𝕜 E' _ _ _ v (f x) - calc - ‖⟪v, f x⟫‖ ≤ ‖v‖ * ‖f x‖ := h₂ - _ ≤ ‖v‖ * (‖f‖ * ‖x‖) := mul_le_mul_of_nonneg_left h₁ (norm_nonneg v) - _ = ‖f‖ * ‖v‖ * ‖x‖ := by ring - - -end ContinuousLinearMap - -section - -variable {ι : Type*} {ι' : Type*} {E' : Type*} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] - -@[simp] -theorem Orthonormal.equiv_refl {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) : - hv.equiv hv (Equiv.refl ι) = LinearIsometryEquiv.refl 𝕜 E := - v.ext_linearIsometryEquiv fun i => by - simp only [Orthonormal.equiv_apply, Equiv.coe_refl, id, LinearIsometryEquiv.coe_refl] - -@[simp] -theorem Orthonormal.equiv_symm {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) {v' : Basis ι' 𝕜 E'} - (hv' : Orthonormal 𝕜 v') (e : ι ≃ ι') : (hv.equiv hv' e).symm = hv'.equiv hv e.symm := - v'.ext_linearIsometryEquiv fun i => - (hv.equiv hv' e).injective <| by - simp only [LinearIsometryEquiv.apply_symm_apply, Orthonormal.equiv_apply, e.apply_symm_apply] - -end - -variable (𝕜) - -/-- `innerSL` is an isometry. Note that the associated `LinearIsometry` is defined in -`InnerProductSpace.Dual` as `toDualMap`. -/ -@[simp] -theorem innerSL_apply_norm (x : E) : ‖innerSL 𝕜 x‖ = ‖x‖ := by - refine - le_antisymm ((innerSL 𝕜 x).opNorm_le_bound (norm_nonneg _) fun y => norm_inner_le_norm _ _) ?_ - rcases (norm_nonneg x).eq_or_gt with (h | h) - · simp [h] - · refine (mul_le_mul_right h).mp ?_ - calc - ‖x‖ * ‖x‖ = ‖(⟪x, x⟫ : 𝕜)‖ := by - rw [← sq, inner_self_eq_norm_sq_to_K, norm_pow, norm_ofReal, abs_norm] - _ ≤ ‖innerSL 𝕜 x‖ * ‖x‖ := (innerSL 𝕜 x).le_opNorm _ - -lemma norm_innerSL_le : ‖innerSL 𝕜 (E := E)‖ ≤ 1 := - ContinuousLinearMap.opNorm_le_bound _ zero_le_one (by simp) - -variable {𝕜} - /-- When an inner product space `E` over `𝕜` is considered as a real normed space, its inner product satisfies `IsBoundedBilinearMap`. @@ -1685,10 +640,6 @@ variable {ι : Type*} local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y -local notation "IK" => @RCLike.I 𝕜 _ - -local postfix:90 "†" => starRingEnd _ - /-- Formula for the distance between the images of two nonzero points under an inversion with center zero. See also `EuclideanGeometry.dist_inversion_inversion` for inversions around a general point. -/ @@ -1882,57 +833,6 @@ theorem eq_of_norm_le_re_inner_eq_norm_sq {x y : E} (hle : ‖x‖ ≤ ‖y‖) end Norm -section BesselsInequality - -variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] - -variable {ι : Type*} (x : E) {v : ι → E} - -local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y - -local notation "IK" => @RCLike.I 𝕜 _ - -local postfix:90 "†" => starRingEnd _ - -/-- Bessel's inequality for finite sums. -/ -theorem Orthonormal.sum_inner_products_le {s : Finset ι} (hv : Orthonormal 𝕜 v) : - ∑ i ∈ s, ‖⟪v i, x⟫‖ ^ 2 ≤ ‖x‖ ^ 2 := by - have h₂ : - (∑ i ∈ s, ∑ j ∈ s, ⟪v i, x⟫ * ⟪x, v j⟫ * ⟪v j, v i⟫) = (∑ k ∈ s, ⟪v k, x⟫ * ⟪x, v k⟫ : 𝕜) := by - classical exact hv.inner_left_right_finset - have h₃ : ∀ z : 𝕜, re (z * conj z) = ‖z‖ ^ 2 := by - intro z - simp only [mul_conj, normSq_eq_def'] - norm_cast - suffices hbf : ‖x - ∑ i ∈ s, ⟪v i, x⟫ • v i‖ ^ 2 = ‖x‖ ^ 2 - ∑ i ∈ s, ‖⟪v i, x⟫‖ ^ 2 by - rw [← sub_nonneg, ← hbf] - simp only [norm_nonneg, pow_nonneg] - rw [@norm_sub_sq 𝕜, sub_add] - simp only [@InnerProductSpace.norm_sq_eq_inner 𝕜, inner_sum, sum_inner] - simp only [inner_smul_right, two_mul, inner_smul_left, inner_conj_symm, ← mul_assoc, h₂, - add_sub_cancel_right, sub_right_inj] - simp only [map_sum, ← inner_conj_symm x, ← h₃] - -/-- Bessel's inequality. -/ -theorem Orthonormal.tsum_inner_products_le (hv : Orthonormal 𝕜 v) : - ∑' i, ‖⟪v i, x⟫‖ ^ 2 ≤ ‖x‖ ^ 2 := by - refine tsum_le_of_sum_le' ?_ fun s => hv.sum_inner_products_le x - simp only [norm_nonneg, pow_nonneg] - -/-- The sum defined in Bessel's inequality is summable. -/ -theorem Orthonormal.inner_products_summable (hv : Orthonormal 𝕜 v) : - Summable fun i => ‖⟪v i, x⟫‖ ^ 2 := by - use ⨆ s : Finset ι, ∑ i ∈ s, ‖⟪v i, x⟫‖ ^ 2 - apply hasSum_of_isLUB_of_nonneg - · intro b - simp only [norm_nonneg, pow_nonneg] - · refine isLUB_ciSup ?_ - use ‖x‖ ^ 2 - rintro y ⟨s, rfl⟩ - exact hv.sum_inner_products_le x - -end BesselsInequality - section RCLike local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y @@ -1951,272 +851,6 @@ theorem RCLike.inner_apply (x y : 𝕜) : ⟪x, y⟫ = conj x * y := end RCLike -section Submodule - -variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] - -local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y - -local notation "IK" => @RCLike.I 𝕜 _ - -local postfix:90 "†" => starRingEnd _ - -/-! ### Inner product space structure on subspaces -/ - -/-- Induced inner product on a submodule. -/ -instance Submodule.innerProductSpace (W : Submodule 𝕜 E) : InnerProductSpace 𝕜 W := - { Submodule.normedSpace W with - inner := fun x y => ⟪(x : E), (y : E)⟫ - conj_symm := fun _ _ => inner_conj_symm _ _ - norm_sq_eq_inner := fun x => norm_sq_eq_inner (x : E) - add_left := fun _ _ _ => inner_add_left _ _ _ - smul_left := fun _ _ _ => inner_smul_left _ _ _ } - -/-- The inner product on submodules is the same as on the ambient space. -/ -@[simp] -theorem Submodule.coe_inner (W : Submodule 𝕜 E) (x y : W) : ⟪x, y⟫ = ⟪(x : E), ↑y⟫ := - rfl - -theorem Orthonormal.codRestrict {ι : Type*} {v : ι → E} (hv : Orthonormal 𝕜 v) (s : Submodule 𝕜 E) - (hvs : ∀ i, v i ∈ s) : @Orthonormal 𝕜 s _ _ _ ι (Set.codRestrict v s hvs) := - s.subtypeₗᵢ.orthonormal_comp_iff.mp hv - -theorem orthonormal_span {ι : Type*} {v : ι → E} (hv : Orthonormal 𝕜 v) : - @Orthonormal 𝕜 (Submodule.span 𝕜 (Set.range v)) _ _ _ ι fun i : ι => - ⟨v i, Submodule.subset_span (Set.mem_range_self i)⟩ := - hv.codRestrict (Submodule.span 𝕜 (Set.range v)) fun i => - Submodule.subset_span (Set.mem_range_self i) - -end Submodule - -/-! ### Families of mutually-orthogonal subspaces of an inner product space -/ - -section OrthogonalFamily_Seminormed - -variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] - -local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y - -local notation "IK" => @RCLike.I 𝕜 _ - -local postfix:90 "†" => starRingEnd _ - -variable {ι : Type*} (𝕜) - -open DirectSum - -/-- An indexed family of mutually-orthogonal subspaces of an inner product space `E`. - -The simple way to express this concept would be as a condition on `V : ι → Submodule 𝕜 E`. We -instead implement it as a condition on a family of inner product spaces each equipped with an -isometric embedding into `E`, thus making it a property of morphisms rather than subobjects. -The connection to the subobject spelling is shown in `orthogonalFamily_iff_pairwise`. - -This definition is less lightweight, but allows for better definitional properties when the inner -product space structure on each of the submodules is important -- for example, when considering -their Hilbert sum (`PiLp V 2`). For example, given an orthonormal set of vectors `v : ι → E`, -we have an associated orthogonal family of one-dimensional subspaces of `E`, which it is convenient -to be able to discuss using `ι → 𝕜` rather than `Π i : ι, span 𝕜 (v i)`. -/ -def OrthogonalFamily (G : ι → Type*) [∀ i, SeminormedAddCommGroup (G i)] - [∀ i, InnerProductSpace 𝕜 (G i)] (V : ∀ i, G i →ₗᵢ[𝕜] E) : Prop := - Pairwise fun i j => ∀ v : G i, ∀ w : G j, ⟪V i v, V j w⟫ = 0 - -variable {𝕜} -variable {G : ι → Type*} [∀ i, NormedAddCommGroup (G i)] [∀ i, InnerProductSpace 𝕜 (G i)] - {V : ∀ i, G i →ₗᵢ[𝕜] E} - -theorem Orthonormal.orthogonalFamily {v : ι → E} (hv : Orthonormal 𝕜 v) : - OrthogonalFamily 𝕜 (fun _i : ι => 𝕜) fun i => LinearIsometry.toSpanSingleton 𝕜 E (hv.1 i) := - fun i j hij a b => by simp [inner_smul_left, inner_smul_right, hv.2 hij] - -section -variable (hV : OrthogonalFamily 𝕜 G V) -include hV - -theorem OrthogonalFamily.eq_ite [DecidableEq ι] {i j : ι} (v : G i) (w : G j) : - ⟪V i v, V j w⟫ = ite (i = j) ⟪V i v, V j w⟫ 0 := by - split_ifs with h - · rfl - · exact hV h v w - -theorem OrthogonalFamily.inner_right_dfinsupp - [∀ (i) (x : G i), Decidable (x ≠ 0)] [DecidableEq ι] (l : ⨁ i, G i) (i : ι) (v : G i) : - ⟪V i v, l.sum fun j => V j⟫ = ⟪v, l i⟫ := - calc - ⟪V i v, l.sum fun j => V j⟫ = l.sum fun j => fun w => ⟪V i v, V j w⟫ := - DFinsupp.inner_sum (fun j => V j) l (V i v) - _ = l.sum fun j => fun w => ite (i = j) ⟪V i v, V j w⟫ 0 := - (congr_arg l.sum <| funext fun _ => funext <| hV.eq_ite v) - _ = ⟪v, l i⟫ := by - simp only [DFinsupp.sum, Submodule.coe_inner, Finset.sum_ite_eq, ite_eq_left_iff, - DFinsupp.mem_support_toFun] - split_ifs with h - · simp only [LinearIsometry.inner_map_map] - · simp only [of_not_not h, inner_zero_right] - -theorem OrthogonalFamily.inner_right_fintype [Fintype ι] (l : ∀ i, G i) (i : ι) (v : G i) : - ⟪V i v, ∑ j : ι, V j (l j)⟫ = ⟪v, l i⟫ := by - classical - calc - ⟪V i v, ∑ j : ι, V j (l j)⟫ = ∑ j : ι, ⟪V i v, V j (l j)⟫ := by rw [inner_sum] - _ = ∑ j, ite (i = j) ⟪V i v, V j (l j)⟫ 0 := - (congr_arg (Finset.sum Finset.univ) <| funext fun j => hV.eq_ite v (l j)) - _ = ⟪v, l i⟫ := by - simp only [Finset.sum_ite_eq, Finset.mem_univ, (V i).inner_map_map, if_true] - -nonrec theorem OrthogonalFamily.inner_sum (l₁ l₂ : ∀ i, G i) (s : Finset ι) : - ⟪∑ i ∈ s, V i (l₁ i), ∑ j ∈ s, V j (l₂ j)⟫ = ∑ i ∈ s, ⟪l₁ i, l₂ i⟫ := by - classical - calc - ⟪∑ i ∈ s, V i (l₁ i), ∑ j ∈ s, V j (l₂ j)⟫ = ∑ j ∈ s, ∑ i ∈ s, ⟪V i (l₁ i), V j (l₂ j)⟫ := by - simp only [sum_inner, inner_sum] - _ = ∑ j ∈ s, ∑ i ∈ s, ite (i = j) ⟪V i (l₁ i), V j (l₂ j)⟫ 0 := by - congr with i - congr with j - apply hV.eq_ite - _ = ∑ i ∈ s, ⟪l₁ i, l₂ i⟫ := by - simp only [Finset.sum_ite_of_true, Finset.sum_ite_eq', LinearIsometry.inner_map_map, - imp_self, imp_true_iff] - -theorem OrthogonalFamily.norm_sum (l : ∀ i, G i) (s : Finset ι) : - ‖∑ i ∈ s, V i (l i)‖ ^ 2 = ∑ i ∈ s, ‖l i‖ ^ 2 := by - have : ((‖∑ i ∈ s, V i (l i)‖ : ℝ) : 𝕜) ^ 2 = ∑ i ∈ s, ((‖l i‖ : ℝ) : 𝕜) ^ 2 := by - simp only [← inner_self_eq_norm_sq_to_K, hV.inner_sum] - exact mod_cast this - -/-- The composition of an orthogonal family of subspaces with an injective function is also an -orthogonal family. -/ -theorem OrthogonalFamily.comp {γ : Type*} {f : γ → ι} (hf : Function.Injective f) : - OrthogonalFamily 𝕜 (fun g => G (f g)) fun g => V (f g) := - fun _i _j hij v w => hV (hf.ne hij) v w - -theorem OrthogonalFamily.orthonormal_sigma_orthonormal {α : ι → Type*} {v_family : ∀ i, α i → G i} - (hv_family : ∀ i, Orthonormal 𝕜 (v_family i)) : - Orthonormal 𝕜 fun a : Σi, α i => V a.1 (v_family a.1 a.2) := by - constructor - · rintro ⟨i, v⟩ - simpa only [LinearIsometry.norm_map] using (hv_family i).left v - rintro ⟨i, v⟩ ⟨j, w⟩ hvw - by_cases hij : i = j - · subst hij - have : v ≠ w := fun h => by - subst h - exact hvw rfl - simpa only [LinearIsometry.inner_map_map] using (hv_family i).2 this - · exact hV hij (v_family i v) (v_family j w) - -theorem OrthogonalFamily.norm_sq_diff_sum [DecidableEq ι] (f : ∀ i, G i) (s₁ s₂ : Finset ι) : - ‖(∑ i ∈ s₁, V i (f i)) - ∑ i ∈ s₂, V i (f i)‖ ^ 2 = - (∑ i ∈ s₁ \ s₂, ‖f i‖ ^ 2) + ∑ i ∈ s₂ \ s₁, ‖f i‖ ^ 2 := by - rw [← Finset.sum_sdiff_sub_sum_sdiff, sub_eq_add_neg, ← Finset.sum_neg_distrib] - let F : ∀ i, G i := fun i => if i ∈ s₁ then f i else -f i - have hF₁ : ∀ i ∈ s₁ \ s₂, F i = f i := fun i hi => if_pos (Finset.sdiff_subset hi) - have hF₂ : ∀ i ∈ s₂ \ s₁, F i = -f i := fun i hi => if_neg (Finset.mem_sdiff.mp hi).2 - have hF : ∀ i, ‖F i‖ = ‖f i‖ := by - intro i - dsimp only [F] - split_ifs <;> simp only [eq_self_iff_true, norm_neg] - have : - ‖(∑ i ∈ s₁ \ s₂, V i (F i)) + ∑ i ∈ s₂ \ s₁, V i (F i)‖ ^ 2 = - (∑ i ∈ s₁ \ s₂, ‖F i‖ ^ 2) + ∑ i ∈ s₂ \ s₁, ‖F i‖ ^ 2 := by - have hs : Disjoint (s₁ \ s₂) (s₂ \ s₁) := disjoint_sdiff_sdiff - simpa only [Finset.sum_union hs] using hV.norm_sum F (s₁ \ s₂ ∪ s₂ \ s₁) - convert this using 4 - · refine Finset.sum_congr rfl fun i hi => ?_ - simp only [hF₁ i hi] - · refine Finset.sum_congr rfl fun i hi => ?_ - simp only [hF₂ i hi, LinearIsometry.map_neg] - · simp only [hF] - · simp only [hF] - -/-- A family `f` of mutually-orthogonal elements of `E` is summable, if and only if -`(fun i ↦ ‖f i‖ ^ 2)` is summable. -/ -theorem OrthogonalFamily.summable_iff_norm_sq_summable [CompleteSpace E] (f : ∀ i, G i) : - (Summable fun i => V i (f i)) ↔ Summable fun i => ‖f i‖ ^ 2 := by - classical - simp only [summable_iff_cauchySeq_finset, NormedAddCommGroup.cauchySeq_iff, Real.norm_eq_abs] - constructor - · intro hf ε hε - obtain ⟨a, H⟩ := hf _ (sqrt_pos.mpr hε) - use a - intro s₁ hs₁ s₂ hs₂ - rw [← Finset.sum_sdiff_sub_sum_sdiff] - refine (abs_sub _ _).trans_lt ?_ - have : ∀ i, 0 ≤ ‖f i‖ ^ 2 := fun i : ι => sq_nonneg _ - simp only [Finset.abs_sum_of_nonneg' this] - have : ((∑ i ∈ s₁ \ s₂, ‖f i‖ ^ 2) + ∑ i ∈ s₂ \ s₁, ‖f i‖ ^ 2) < √ε ^ 2 := by - rw [← hV.norm_sq_diff_sum, sq_lt_sq, abs_of_nonneg (sqrt_nonneg _), - abs_of_nonneg (norm_nonneg _)] - exact H s₁ hs₁ s₂ hs₂ - have hη := sq_sqrt (le_of_lt hε) - linarith - · intro hf ε hε - have hε' : 0 < ε ^ 2 / 2 := half_pos (sq_pos_of_pos hε) - obtain ⟨a, H⟩ := hf _ hε' - use a - intro s₁ hs₁ s₂ hs₂ - refine (abs_lt_of_sq_lt_sq' ?_ (le_of_lt hε)).2 - have has : a ≤ s₁ ⊓ s₂ := le_inf hs₁ hs₂ - rw [hV.norm_sq_diff_sum] - have Hs₁ : ∑ x ∈ s₁ \ s₂, ‖f x‖ ^ 2 < ε ^ 2 / 2 := by - convert H _ hs₁ _ has - have : s₁ ⊓ s₂ ⊆ s₁ := Finset.inter_subset_left - rw [← Finset.sum_sdiff this, add_tsub_cancel_right, Finset.abs_sum_of_nonneg'] - · simp - · exact fun i => sq_nonneg _ - have Hs₂ : ∑ x ∈ s₂ \ s₁, ‖f x‖ ^ 2 < ε ^ 2 / 2 := by - convert H _ hs₂ _ has - have : s₁ ⊓ s₂ ⊆ s₂ := Finset.inter_subset_right - rw [← Finset.sum_sdiff this, add_tsub_cancel_right, Finset.abs_sum_of_nonneg'] - · simp - · exact fun i => sq_nonneg _ - linarith - -end - -end OrthogonalFamily_Seminormed - -section OrthogonalFamily - -variable [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] - -local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y - -local notation "IK" => @RCLike.I 𝕜 _ - -local postfix:90 "†" => starRingEnd _ - -variable {ι : Type*} {G : ι → Type*} - -/-- An orthogonal family forms an independent family of subspaces; that is, any collection of -elements each from a different subspace in the family is linearly independent. In particular, the -pairwise intersections of elements of the family are 0. -/ -theorem OrthogonalFamily.independent {V : ι → Submodule 𝕜 E} - (hV : OrthogonalFamily 𝕜 (fun i => V i) fun i => (V i).subtypeₗᵢ) : - iSupIndep V := by - classical - apply iSupIndep_of_dfinsupp_lsum_injective - refine LinearMap.ker_eq_bot.mp ?_ - rw [Submodule.eq_bot_iff] - intro v hv - rw [LinearMap.mem_ker] at hv - ext i - suffices ⟪(v i : E), v i⟫ = 0 by simpa only [inner_self_eq_zero] using this - calc - ⟪(v i : E), v i⟫ = ⟪(v i : E), DFinsupp.lsum ℕ (fun i => (V i).subtype) v⟫ := by - simpa only [DFinsupp.sumAddHom_apply, DFinsupp.lsum_apply_apply] using - (hV.inner_right_dfinsupp v i (v i)).symm - _ = 0 := by simp only [hv, inner_zero_right] - -theorem DirectSum.IsInternal.collectedBasis_orthonormal [DecidableEq ι] {V : ι → Submodule 𝕜 E} - (hV : OrthogonalFamily 𝕜 (fun i => V i) fun i => (V i).subtypeₗᵢ) - (hV_sum : DirectSum.IsInternal fun i => V i) {α : ι → Type*} - {v_family : ∀ i, Basis (α i) 𝕜 (V i)} (hv_family : ∀ i, Orthonormal 𝕜 (v_family i)) : - Orthonormal 𝕜 (hV_sum.collectedBasis v_family) := by - simpa only [hV_sum.collectedBasis_coe] using hV.orthonormal_sigma_orthonormal hv_family - -end OrthogonalFamily - section RCLikeToReal open scoped InnerProductSpace @@ -2227,10 +861,6 @@ variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y -local notation "IK" => @RCLike.I 𝕜 _ - -local postfix:90 "†" => starRingEnd _ - /-- A general inner product implies a real inner product. This is not registered as an instance since it creates problems with the case `𝕜 = ℝ`. -/ def Inner.rclikeToReal : Inner ℝ E where inner x y := re ⟪x, y⟫ @@ -2275,13 +905,6 @@ instance : InnerProductSpace ℝ ℂ := InnerProductSpace.complexToReal protected theorem Complex.inner (w z : ℂ) : ⟪w, z⟫_ℝ = (conj w * z).re := rfl -/-- The inner product on an inner product space of dimension 2 can be evaluated in terms -of a complex-number representation of the space. -/ -theorem inner_map_complex [SeminormedAddCommGroup G] [InnerProductSpace ℝ G] (f : G ≃ₗᵢ[ℝ] ℂ) - (x y : G) : ⟪x, y⟫_ℝ = (conj (f x) * f y).re := by rw [← Complex.inner, f.inner_map_map] - -instance : InnerProductSpace ℝ ℂ := InnerProductSpace.complexToReal - end RCLikeToReal section Continuous @@ -2290,15 +913,10 @@ variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y -local notation "IK" => @RCLike.I 𝕜 _ - -local postfix:90 "†" => starRingEnd _ - /-! ### Continuity of the inner product -/ - theorem continuous_inner : Continuous fun p : E × E => ⟪p.1, p.2⟫ := letI : InnerProductSpace ℝ E := InnerProductSpace.rclikeToReal 𝕜 E letI : IsScalarTower ℝ 𝕜 E := RestrictScalars.isScalarTower _ _ _ @@ -2328,146 +946,3 @@ theorem Continuous.inner (hf : Continuous f) (hg : Continuous g) : Continuous fu continuous_iff_continuousAt.2 fun _x => hf.continuousAt.inner hg.continuousAt end Continuous - -section ReApplyInnerSelf - -variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] - -local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y - -local notation "IK" => @RCLike.I 𝕜 _ - -local postfix:90 "†" => starRingEnd _ - -/-- Extract a real bilinear form from an operator `T`, -by taking the pairing `fun x ↦ re ⟪T x, x⟫`. -/ -def ContinuousLinearMap.reApplyInnerSelf (T : E →L[𝕜] E) (x : E) : ℝ := - re ⟪T x, x⟫ - -theorem ContinuousLinearMap.reApplyInnerSelf_apply (T : E →L[𝕜] E) (x : E) : - T.reApplyInnerSelf x = re ⟪T x, x⟫ := - rfl - -end ReApplyInnerSelf - -section ReApplyInnerSelf_Seminormed - -variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] - -local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y - -local notation "IK" => @RCLike.I 𝕜 _ - -local postfix:90 "†" => starRingEnd _ - -theorem ContinuousLinearMap.reApplyInnerSelf_continuous (T : E →L[𝕜] E) : - Continuous T.reApplyInnerSelf := - reCLM.continuous.comp <| T.continuous.inner continuous_id - -theorem ContinuousLinearMap.reApplyInnerSelf_smul (T : E →L[𝕜] E) (x : E) {c : 𝕜} : - T.reApplyInnerSelf (c • x) = ‖c‖ ^ 2 * T.reApplyInnerSelf x := by - simp only [ContinuousLinearMap.map_smul, ContinuousLinearMap.reApplyInnerSelf_apply, - inner_smul_left, inner_smul_right, ← mul_assoc, mul_conj, ← ofReal_pow, ← smul_re, - Algebra.smul_def (‖c‖ ^ 2) ⟪T x, x⟫, algebraMap_eq_ofReal] - -end ReApplyInnerSelf_Seminormed - -section SeparationQuotient -variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] - -theorem Inseparable.inner_eq_inner {x₁ x₂ y₁ y₂ : E} - (hx : Inseparable x₁ x₂) (hy : Inseparable y₁ y₂) : - inner x₁ y₁ = (inner x₂ y₂ : 𝕜) := - ((hx.prod hy).map continuous_inner).eq - -namespace SeparationQuotient - -instance : Inner 𝕜 (SeparationQuotient E) where - inner := SeparationQuotient.lift₂ Inner.inner fun _ _ _ _ => Inseparable.inner_eq_inner - -@[simp] -theorem inner_mk_mk (x y : E) : - inner (mk x) (mk y) = (inner x y : 𝕜) := rfl - -instance : InnerProductSpace 𝕜 (SeparationQuotient E) where - norm_sq_eq_inner := Quotient.ind norm_sq_eq_inner - conj_symm := Quotient.ind₂ inner_conj_symm - add_left := Quotient.ind fun x => Quotient.ind₂ <| inner_add_left x - smul_left := Quotient.ind₂ inner_smul_left - -end SeparationQuotient - -end SeparationQuotient - -section UniformSpace.Completion - -variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] - -local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y - -local notation "IK" => @RCLike.I 𝕜 _ - -local postfix:90 "†" => starRingEnd _ - -namespace UniformSpace.Completion - -open UniformSpace Function - -instance toInner {𝕜' E' : Type*} [TopologicalSpace 𝕜'] [UniformSpace E'] [Inner 𝕜' E'] : - Inner 𝕜' (Completion E') where - inner := curry <| (isDenseInducing_coe.prodMap isDenseInducing_coe).extend (uncurry inner) - -@[simp] -theorem inner_coe (a b : E) : inner (a : Completion E) (b : Completion E) = (inner a b : 𝕜) := - (isDenseInducing_coe.prodMap isDenseInducing_coe).extend_eq - (continuous_inner : Continuous (uncurry inner : E × E → 𝕜)) (a, b) - -protected theorem continuous_inner : - Continuous (uncurry inner : Completion E × Completion E → 𝕜) := by - let inner' : E →+ E →+ 𝕜 := - { toFun := fun x => (innerₛₗ 𝕜 x).toAddMonoidHom - map_zero' := by ext x; exact inner_zero_left _ - map_add' := fun x y => by ext z; exact inner_add_left _ _ _ } - have : Continuous fun p : E × E => inner' p.1 p.2 := continuous_inner - rw [Completion.toInner, inner, uncurry_curry _] - change - Continuous - (((isDenseInducing_toCompl E).prodMap (isDenseInducing_toCompl E)).extend fun p : E × E => - inner' p.1 p.2) - exact (isDenseInducing_toCompl E).extend_Z_bilin (isDenseInducing_toCompl E) this - -protected theorem Continuous.inner {α : Type*} [TopologicalSpace α] {f g : α → Completion E} - (hf : Continuous f) (hg : Continuous g) : Continuous (fun x : α => inner (f x) (g x) : α → 𝕜) := - UniformSpace.Completion.continuous_inner.comp (hf.prod_mk hg : _) - -instance innerProductSpace : InnerProductSpace 𝕜 (Completion E) where - norm_sq_eq_inner x := - Completion.induction_on x - (isClosed_eq (continuous_norm.pow 2) - (continuous_re.comp (Continuous.inner continuous_id' continuous_id'))) - fun a => by simp only [norm_coe, inner_coe, inner_self_eq_norm_sq] - conj_symm x y := - Completion.induction_on₂ x y - (isClosed_eq (continuous_conj.comp (Continuous.inner continuous_snd continuous_fst)) - (Continuous.inner continuous_fst continuous_snd)) - fun a b => by simp only [inner_coe, inner_conj_symm] - add_left x y z := - Completion.induction_on₃ x y z - (isClosed_eq - (Continuous.inner (continuous_fst.add (continuous_fst.comp continuous_snd)) - (continuous_snd.comp continuous_snd)) - ((Continuous.inner continuous_fst (continuous_snd.comp continuous_snd)).add - (Continuous.inner (continuous_fst.comp continuous_snd) - (continuous_snd.comp continuous_snd)))) - fun a b c => by simp only [← coe_add, inner_coe, inner_add_left] - smul_left x y c := - Completion.induction_on₂ x y - (isClosed_eq (Continuous.inner (continuous_fst.const_smul c) continuous_snd) - ((continuous_mul_left _).comp (Continuous.inner continuous_fst continuous_snd))) - fun a b => by simp only [← coe_smul c a, inner_coe, inner_smul_left] - -end UniformSpace.Completion - -end UniformSpace.Completion - -set_option linter.style.longFile 2500 diff --git a/Mathlib/Analysis/InnerProductSpace/Completion.lean b/Mathlib/Analysis/InnerProductSpace/Completion.lean new file mode 100644 index 0000000000000..330ff19ec4db2 --- /dev/null +++ b/Mathlib/Analysis/InnerProductSpace/Completion.lean @@ -0,0 +1,119 @@ +/- +Copyright (c) 2019 Zhouhang Zhou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Zhouhang Zhou, Sébastien Gouëzel, Frédéric Dupuis +-/ +import Mathlib.Analysis.InnerProductSpace.LinearMap +import Mathlib.Analysis.Normed.Module.Completion + +/-! +# Completion of an inner product space + +We show that the separation quotient and the completion of an inner product space are inner +product spaces. +-/ + +noncomputable section + +open RCLike Real Filter Topology ComplexConjugate Finsupp +open LinearMap (BilinForm) + +variable {𝕜 E F : Type*} [RCLike 𝕜] + +section SeparationQuotient +variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] + +theorem Inseparable.inner_eq_inner {x₁ x₂ y₁ y₂ : E} + (hx : Inseparable x₁ x₂) (hy : Inseparable y₁ y₂) : + inner x₁ y₁ = (inner x₂ y₂ : 𝕜) := + ((hx.prod hy).map continuous_inner).eq + +namespace SeparationQuotient + +instance : Inner 𝕜 (SeparationQuotient E) where + inner := SeparationQuotient.lift₂ Inner.inner fun _ _ _ _ => Inseparable.inner_eq_inner + +@[simp] +theorem inner_mk_mk (x y : E) : + inner (mk x) (mk y) = (inner x y : 𝕜) := rfl + +instance : InnerProductSpace 𝕜 (SeparationQuotient E) where + norm_sq_eq_inner := Quotient.ind norm_sq_eq_inner + conj_symm := Quotient.ind₂ inner_conj_symm + add_left := Quotient.ind fun x => Quotient.ind₂ <| inner_add_left x + smul_left := Quotient.ind₂ inner_smul_left + +end SeparationQuotient + +end SeparationQuotient + +section UniformSpace.Completion + +variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + +local notation "IK" => @RCLike.I 𝕜 _ + +local postfix:90 "†" => starRingEnd _ + +namespace UniformSpace.Completion + +open UniformSpace Function + +instance toInner {𝕜' E' : Type*} [TopologicalSpace 𝕜'] [UniformSpace E'] [Inner 𝕜' E'] : + Inner 𝕜' (Completion E') where + inner := curry <| (isDenseInducing_coe.prodMap isDenseInducing_coe).extend (uncurry inner) + +@[simp] +theorem inner_coe (a b : E) : inner (a : Completion E) (b : Completion E) = (inner a b : 𝕜) := + (isDenseInducing_coe.prodMap isDenseInducing_coe).extend_eq + (continuous_inner : Continuous (uncurry inner : E × E → 𝕜)) (a, b) + +protected theorem continuous_inner : + Continuous (uncurry inner : Completion E × Completion E → 𝕜) := by + let inner' : E →+ E →+ 𝕜 := + { toFun := fun x => (innerₛₗ 𝕜 x).toAddMonoidHom + map_zero' := by ext x; exact inner_zero_left _ + map_add' := fun x y => by ext z; exact inner_add_left _ _ _ } + have : Continuous fun p : E × E => inner' p.1 p.2 := continuous_inner + rw [Completion.toInner, inner, uncurry_curry _] + change + Continuous + (((isDenseInducing_toCompl E).prodMap (isDenseInducing_toCompl E)).extend fun p : E × E => + inner' p.1 p.2) + exact (isDenseInducing_toCompl E).extend_Z_bilin (isDenseInducing_toCompl E) this + +protected theorem Continuous.inner {α : Type*} [TopologicalSpace α] {f g : α → Completion E} + (hf : Continuous f) (hg : Continuous g) : Continuous (fun x : α => inner (f x) (g x) : α → 𝕜) := + UniformSpace.Completion.continuous_inner.comp (hf.prod_mk hg : _) + +instance innerProductSpace : InnerProductSpace 𝕜 (Completion E) where + norm_sq_eq_inner x := + Completion.induction_on x + (isClosed_eq (continuous_norm.pow 2) + (continuous_re.comp (Continuous.inner continuous_id' continuous_id'))) + fun a => by simp only [norm_coe, inner_coe, inner_self_eq_norm_sq] + conj_symm x y := + Completion.induction_on₂ x y + (isClosed_eq (continuous_conj.comp (Continuous.inner continuous_snd continuous_fst)) + (Continuous.inner continuous_fst continuous_snd)) + fun a b => by simp only [inner_coe, inner_conj_symm] + add_left x y z := + Completion.induction_on₃ x y z + (isClosed_eq + (Continuous.inner (continuous_fst.add (continuous_fst.comp continuous_snd)) + (continuous_snd.comp continuous_snd)) + ((Continuous.inner continuous_fst (continuous_snd.comp continuous_snd)).add + (Continuous.inner (continuous_fst.comp continuous_snd) + (continuous_snd.comp continuous_snd)))) + fun a b c => by simp only [← coe_add, inner_coe, inner_add_left] + smul_left x y c := + Completion.induction_on₂ x y + (isClosed_eq (Continuous.inner (continuous_fst.const_smul c) continuous_snd) + ((continuous_mul_left _).comp (Continuous.inner continuous_fst continuous_snd))) + fun a b => by simp only [← coe_smul c a, inner_coe, inner_smul_left] + +end UniformSpace.Completion + +end UniformSpace.Completion diff --git a/Mathlib/Analysis/InnerProductSpace/ConformalLinearMap.lean b/Mathlib/Analysis/InnerProductSpace/ConformalLinearMap.lean index 1a056881fc625..773370669fb2f 100644 --- a/Mathlib/Analysis/InnerProductSpace/ConformalLinearMap.lean +++ b/Mathlib/Analysis/InnerProductSpace/ConformalLinearMap.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yourong Zang -/ import Mathlib.Analysis.NormedSpace.ConformalLinearMap -import Mathlib.Analysis.InnerProductSpace.Basic +import Mathlib.Analysis.InnerProductSpace.LinearMap /-! # Conformal maps between inner product spaces diff --git a/Mathlib/Analysis/InnerProductSpace/Defs.lean b/Mathlib/Analysis/InnerProductSpace/Defs.lean new file mode 100644 index 0000000000000..c89fcb0886d38 --- /dev/null +++ b/Mathlib/Analysis/InnerProductSpace/Defs.lean @@ -0,0 +1,502 @@ +/- +Copyright (c) 2019 Zhouhang Zhou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Zhouhang Zhou, Sébastien Gouëzel, Frédéric Dupuis +-/ +import Mathlib.Algebra.QuadraticDiscriminant +import Mathlib.Analysis.RCLike.Basic +import Mathlib.Data.Complex.Basic + +/-! +# Inner product spaces + +This file defines inner product spaces. We do not formally define Hilbert spaces, but they can be +obtained using the set of assumptions +`[NormedAddCommGroup E] [InnerProductSpace 𝕜 E] [CompleteSpace E]`. + +An inner product space is a vector space endowed with an inner product. It generalizes the notion of +dot product in `ℝ^n` and provides the means of defining the length of a vector and the angle between +two vectors. In particular vectors `x` and `y` are orthogonal if their inner product equals zero. +We define both the real and complex cases at the same time using the `RCLike` typeclass. + +Rather than defining the norm on an inner product space to be `√(re ⟪x, x⟫)`, we assume that a norm +is given, and add a hypothesis stating that `‖x‖ ^ 2 = re (inner x x)`. This makes it possible to +handle spaces where the norm is equal, but not defeq, to the square root of the +inner product. Defining a norm starting from an inner product is handled via the +`InnerProductSpace.Core` structure. + +This file is intended to contain the minimal amount of machinery needed to define inner product +spaces, and to construct a normed space from an inner product space. Many more general lemmas can +be found in `Analysis.InnerProductSpace.Basic`. For the specific construction of an inner product +structure on `n → 𝕜` for `𝕜 = ℝ` or `ℂ`, see `EuclideanSpace` in +`Analysis.InnerProductSpace.PiL2`. + +## Main results + +- We define the class `InnerProductSpace 𝕜 E` extending `NormedSpace 𝕜 E` with a number of basic + properties, most notably the Cauchy-Schwarz inequality. Here `𝕜` is understood to be either `ℝ` + or `ℂ`, through the `RCLike` typeclass. + +## Notation + +We globally denote the real and complex inner products by `⟪·, ·⟫_ℝ` and `⟪·, ·⟫_ℂ` respectively. +We also provide two notation namespaces: `RealInnerProductSpace`, `ComplexInnerProductSpace`, +which respectively introduce the plain notation `⟪·, ·⟫` for the real and complex inner product. + +## Implementation notes + +We choose the convention that inner products are conjugate linear in the first argument and linear +in the second. + +## Tags + +inner product space, Hilbert space, norm + +## References +* [Clément & Martin, *The Lax-Milgram Theorem. A detailed proof to be formalized in Coq*] +* [Clément & Martin, *A Coq formal proof of the Lax–Milgram theorem*] + +The Coq code is available at the following address: +-/ + + +noncomputable section + +open RCLike Real Filter Topology ComplexConjugate Finsupp + +open LinearMap (BilinForm) + +variable {𝕜 E F : Type*} [RCLike 𝕜] + +/-- Syntactic typeclass for types endowed with an inner product -/ +class Inner (𝕜 E : Type*) where + /-- The inner product function. -/ + inner : E → E → 𝕜 + +export Inner (inner) + +/-- The inner product with values in `𝕜`. -/ +scoped[InnerProductSpace] notation3:max "⟪" x ", " y "⟫_" 𝕜:max => @inner 𝕜 _ _ x y + +section Notations + +/-- The inner product with values in `ℝ`. -/ +scoped[RealInnerProductSpace] notation "⟪" x ", " y "⟫" => @inner ℝ _ _ x y + +/-- The inner product with values in `ℂ`. -/ +scoped[ComplexInnerProductSpace] notation "⟪" x ", " y "⟫" => @inner ℂ _ _ x y + +end Notations + +/-- A (pre) inner product space is a vector space with an additional operation called inner product. +The (semi)norm could be derived from the inner product, instead we require the existence of a +seminorm and the fact that `‖x‖^2 = re ⟪x, x⟫` to be able to put instances on `𝕂` or product spaces. + +Note that `NormedSpace` does not assume that `‖x‖=0` implies `x=0` (it is rather a seminorm). + +To construct a seminorm from an inner product, see `PreInnerProductSpace.ofCore`. +-/ +class InnerProductSpace (𝕜 : Type*) (E : Type*) [RCLike 𝕜] [SeminormedAddCommGroup E] extends + NormedSpace 𝕜 E, Inner 𝕜 E where + /-- The inner product induces the norm. -/ + norm_sq_eq_inner : ∀ x : E, ‖x‖ ^ 2 = re (inner x x) + /-- The inner product is *hermitian*, taking the `conj` swaps the arguments. -/ + conj_symm : ∀ x y, conj (inner y x) = inner x y + /-- The inner product is additive in the first coordinate. -/ + add_left : ∀ x y z, inner (x + y) z = inner x z + inner y z + /-- The inner product is conjugate linear in the first coordinate. -/ + smul_left : ∀ x y r, inner (r • x) y = conj r * inner x y + +/-! +### Constructing a normed space structure from an inner product + +In the definition of an inner product space, we require the existence of a norm, which is equal +(but maybe not defeq) to the square root of the scalar product. This makes it possible to put +an inner product space structure on spaces with a preexisting norm (for instance `ℝ`), with good +properties. However, sometimes, one would like to define the norm starting only from a well-behaved +scalar product. This is what we implement in this paragraph, starting from a structure +`InnerProductSpace.Core` stating that we have a nice scalar product. + +Our goal here is not to develop a whole theory with all the supporting API, as this will be done +below for `InnerProductSpace`. Instead, we implement the bare minimum to go as directly as +possible to the construction of the norm and the proof of the triangular inequality. + +Warning: Do not use this `Core` structure if the space you are interested in already has a norm +instance defined on it, otherwise this will create a second non-defeq norm instance! +-/ + +/-- A structure requiring that a scalar product is positive semidefinite and symmetric. -/ +structure PreInnerProductSpace.Core (𝕜 : Type*) (F : Type*) [RCLike 𝕜] [AddCommGroup F] + [Module 𝕜 F] extends Inner 𝕜 F where + /-- The inner product is *hermitian*, taking the `conj` swaps the arguments. -/ + conj_symm : ∀ x y, conj (inner y x) = inner x y + /-- The inner product is positive (semi)definite. -/ + nonneg_re : ∀ x, 0 ≤ re (inner x x) + /-- The inner product is additive in the first coordinate. -/ + add_left : ∀ x y z, inner (x + y) z = inner x z + inner y z + /-- The inner product is conjugate linear in the first coordinate. -/ + smul_left : ∀ x y r, inner (r • x) y = conj r * inner x y + +attribute [class] PreInnerProductSpace.Core + +/-- A structure requiring that a scalar product is positive definite. Some theorems that +require this assumptions are put under section `InnerProductSpace.Core`. -/ +-- @[nolint HasNonemptyInstance] porting note: I don't think we have this linter anymore +structure InnerProductSpace.Core (𝕜 : Type*) (F : Type*) [RCLike 𝕜] [AddCommGroup F] + [Module 𝕜 F] extends PreInnerProductSpace.Core 𝕜 F where + /-- The inner product is positive definite. -/ + definite : ∀ x, inner x x = 0 → x = 0 + +/- We set `InnerProductSpace.Core` to be a class as we will use it as such in the construction +of the normed space structure that it produces. However, all the instances we will use will be +local to this proof. -/ +attribute [class] InnerProductSpace.Core + +instance (𝕜 : Type*) (F : Type*) [RCLike 𝕜] [AddCommGroup F] + [Module 𝕜 F] [cd : InnerProductSpace.Core 𝕜 F] : PreInnerProductSpace.Core 𝕜 F where + inner := cd.inner + conj_symm := cd.conj_symm + nonneg_re := cd.nonneg_re + add_left := cd.add_left + smul_left := cd.smul_left + +/-- Define `PreInnerProductSpace.Core` from `PreInnerProductSpace`. Defined to reuse lemmas about +`PreInnerProductSpace.Core` for `PreInnerProductSpace`s. Note that the `Seminorm` instance provided +by `PreInnerProductSpace.Core.norm` is propositionally but not definitionally equal to the original +norm. -/ +def PreInnerProductSpace.toCore [SeminormedAddCommGroup E] [c : InnerProductSpace 𝕜 E] : + PreInnerProductSpace.Core 𝕜 E := + { c with + nonneg_re := fun x => by + rw [← InnerProductSpace.norm_sq_eq_inner] + apply sq_nonneg } + +/-- Define `InnerProductSpace.Core` from `InnerProductSpace`. Defined to reuse lemmas about +`InnerProductSpace.Core` for `InnerProductSpace`s. Note that the `Norm` instance provided by +`InnerProductSpace.Core.norm` is propositionally but not definitionally equal to the original +norm. -/ +def InnerProductSpace.toCore [NormedAddCommGroup E] [c : InnerProductSpace 𝕜 E] : + InnerProductSpace.Core 𝕜 E := + { c with + nonneg_re := fun x => by + rw [← InnerProductSpace.norm_sq_eq_inner] + apply sq_nonneg + definite := fun x hx => + norm_eq_zero.1 <| pow_eq_zero (n := 2) <| by + rw [InnerProductSpace.norm_sq_eq_inner (𝕜 := 𝕜) x, hx, map_zero] } + +namespace InnerProductSpace.Core + +section PreInnerProductSpace.Core + +variable [AddCommGroup F] [Module 𝕜 F] [c : PreInnerProductSpace.Core 𝕜 F] + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 F _ x y + +local notation "normSqK" => @RCLike.normSq 𝕜 _ + +local notation "reK" => @RCLike.re 𝕜 _ + +local notation "ext_iff" => @RCLike.ext_iff 𝕜 _ + +local postfix:90 "†" => starRingEnd _ + +/-- Inner product defined by the `PreInnerProductSpace.Core` structure. We can't reuse +`PreInnerProductSpace.Core.toInner` because it takes `PreInnerProductSpace.Core` as an explicit +argument. -/ +def toPreInner' : Inner 𝕜 F := + c.toInner + +attribute [local instance] toPreInner' + +/-- The norm squared function for `PreInnerProductSpace.Core` structure. -/ +def normSq (x : F) := + reK ⟪x, x⟫ + +local notation "normSqF" => @normSq 𝕜 F _ _ _ _ + +theorem inner_conj_symm (x y : F) : ⟪y, x⟫† = ⟪x, y⟫ := + c.conj_symm x y + +theorem inner_self_nonneg {x : F} : 0 ≤ re ⟪x, x⟫ := + c.nonneg_re _ + +theorem inner_self_im (x : F) : im ⟪x, x⟫ = 0 := by + rw [← @ofReal_inj 𝕜, im_eq_conj_sub] + simp [inner_conj_symm] + +theorem inner_add_left (x y z : F) : ⟪x + y, z⟫ = ⟪x, z⟫ + ⟪y, z⟫ := + c.add_left _ _ _ + +theorem inner_add_right (x y z : F) : ⟪x, y + z⟫ = ⟪x, y⟫ + ⟪x, z⟫ := by + rw [← inner_conj_symm, inner_add_left, RingHom.map_add]; simp only [inner_conj_symm] + +theorem ofReal_normSq_eq_inner_self (x : F) : (normSqF x : 𝕜) = ⟪x, x⟫ := by + rw [ext_iff] + exact ⟨by simp only [ofReal_re]; rfl, by simp only [inner_self_im, ofReal_im]⟩ + +theorem inner_re_symm (x y : F) : re ⟪x, y⟫ = re ⟪y, x⟫ := by rw [← inner_conj_symm, conj_re] + +theorem inner_im_symm (x y : F) : im ⟪x, y⟫ = -im ⟪y, x⟫ := by rw [← inner_conj_symm, conj_im] + +theorem inner_smul_left (x y : F) {r : 𝕜} : ⟪r • x, y⟫ = r† * ⟪x, y⟫ := + c.smul_left _ _ _ + +theorem inner_smul_right (x y : F) {r : 𝕜} : ⟪x, r • y⟫ = r * ⟪x, y⟫ := by + rw [← inner_conj_symm, inner_smul_left] + simp only [conj_conj, inner_conj_symm, RingHom.map_mul] + +theorem inner_zero_left (x : F) : ⟪0, x⟫ = 0 := by + rw [← zero_smul 𝕜 (0 : F), inner_smul_left] + simp only [zero_mul, RingHom.map_zero] + +theorem inner_zero_right (x : F) : ⟪x, 0⟫ = 0 := by + rw [← inner_conj_symm, inner_zero_left]; simp only [RingHom.map_zero] + +theorem inner_self_of_eq_zero {x : F} : x = 0 → ⟪x, x⟫ = 0 := by + rintro rfl + exact inner_zero_left _ + +theorem normSq_eq_zero_of_eq_zero {x : F} : x = 0 → normSqF x = 0 := by + rintro rfl + simp [normSq, inner_self_of_eq_zero] + +theorem ne_zero_of_inner_self_ne_zero {x : F} : ⟪x, x⟫ ≠ 0 → x ≠ 0 := + mt inner_self_of_eq_zero +theorem inner_self_ofReal_re (x : F) : (re ⟪x, x⟫ : 𝕜) = ⟪x, x⟫ := by + norm_num [ext_iff, inner_self_im] + +theorem norm_inner_symm (x y : F) : ‖⟪x, y⟫‖ = ‖⟪y, x⟫‖ := by rw [← inner_conj_symm, norm_conj] + +theorem inner_neg_left (x y : F) : ⟪-x, y⟫ = -⟪x, y⟫ := by + rw [← neg_one_smul 𝕜 x, inner_smul_left] + simp + +theorem inner_neg_right (x y : F) : ⟪x, -y⟫ = -⟪x, y⟫ := by + rw [← inner_conj_symm, inner_neg_left]; simp only [RingHom.map_neg, inner_conj_symm] + +theorem inner_sub_left (x y z : F) : ⟪x - y, z⟫ = ⟪x, z⟫ - ⟪y, z⟫ := by + simp [sub_eq_add_neg, inner_add_left, inner_neg_left] + +theorem inner_sub_right (x y z : F) : ⟪x, y - z⟫ = ⟪x, y⟫ - ⟪x, z⟫ := by + simp [sub_eq_add_neg, inner_add_right, inner_neg_right] + +theorem inner_mul_symm_re_eq_norm (x y : F) : re (⟪x, y⟫ * ⟪y, x⟫) = ‖⟪x, y⟫ * ⟪y, x⟫‖ := by + rw [← inner_conj_symm, mul_comm] + exact re_eq_norm_of_mul_conj (inner y x) + +/-- Expand `inner (x + y) (x + y)` -/ +theorem inner_add_add_self (x y : F) : ⟪x + y, x + y⟫ = ⟪x, x⟫ + ⟪x, y⟫ + ⟪y, x⟫ + ⟪y, y⟫ := by + simp only [inner_add_left, inner_add_right]; ring + +-- Expand `inner (x - y) (x - y)` +theorem inner_sub_sub_self (x y : F) : ⟪x - y, x - y⟫ = ⟪x, x⟫ - ⟪x, y⟫ - ⟪y, x⟫ + ⟪y, y⟫ := by + simp only [inner_sub_left, inner_sub_right]; ring + +theorem inner_smul_ofReal_left (x y : F) {t : ℝ} : ⟪(t : 𝕜) • x, y⟫ = ⟪x, y⟫ * t := by + rw [inner_smul_left, conj_ofReal, mul_comm] + +theorem inner_smul_ofReal_right (x y : F) {t : ℝ} : ⟪x, (t : 𝕜) • y⟫ = ⟪x, y⟫ * t := by + rw [inner_smul_right, mul_comm] + +theorem re_inner_smul_ofReal_smul_self (x : F) {t : ℝ} : + re ⟪(t : 𝕜) • x, (t : 𝕜) • x⟫ = normSqF x * t * t := by + apply ofReal_injective (K := 𝕜) + simp [inner_self_ofReal_re, inner_smul_ofReal_left, inner_smul_ofReal_right, normSq] + +/-- An auxiliary equality useful to prove the **Cauchy–Schwarz inequality**. Here we use the +standard argument involving the discriminant of quadratic form. -/ +lemma cauchy_schwarz_aux' (x y : F) (t : ℝ) : 0 ≤ normSqF x * t * t + 2 * re ⟪x, y⟫ * t + + normSqF y := by + calc 0 ≤ re ⟪(ofReal t : 𝕜) • x + y, (ofReal t : 𝕜) • x + y⟫ := inner_self_nonneg + _ = re (⟪(ofReal t : 𝕜) • x, (ofReal t : 𝕜) • x⟫ + ⟪(ofReal t : 𝕜) • x, y⟫ + + ⟪y, (ofReal t : 𝕜) • x⟫ + ⟪y, y⟫) := by rw [inner_add_add_self ((ofReal t : 𝕜) • x) y] + _ = re ⟪(ofReal t : 𝕜) • x, (ofReal t : 𝕜) • x⟫ + + re ⟪(ofReal t : 𝕜) • x, y⟫ + re ⟪y, (ofReal t : 𝕜) • x⟫ + re ⟪y, y⟫ := by + simp only [map_add] + _ = normSq x * t * t + re (⟪x, y⟫ * t) + re (⟪y, x⟫ * t) + re ⟪y, y⟫ := by rw + [re_inner_smul_ofReal_smul_self, inner_smul_ofReal_left, inner_smul_ofReal_right] + _ = normSq x * t * t + re ⟪x, y⟫ * t + re ⟪y, x⟫ * t + re ⟪y, y⟫ := by rw [mul_comm ⟪x,y⟫ _, + RCLike.re_ofReal_mul, mul_comm t _, mul_comm ⟪y,x⟫ _, RCLike.re_ofReal_mul, mul_comm t _] + _ = normSq x * t * t + re ⟪x, y⟫ * t + re ⟪y, x⟫ * t + normSq y := by rw [← normSq] + _ = normSq x * t * t + re ⟪x, y⟫ * t + re ⟪x, y⟫ * t + normSq y := by rw [inner_re_symm] + _ = normSq x * t * t + 2 * re ⟪x, y⟫ * t + normSq y := by ring + +/-- Another auxiliary equality related with the **Cauchy–Schwarz inequality**: the square of the +seminorm of `⟪x, y⟫ • x - ⟪x, x⟫ • y` is equal to `‖x‖ ^ 2 * (‖x‖ ^ 2 * ‖y‖ ^ 2 - ‖⟪x, y⟫‖ ^ 2)`. +We use `InnerProductSpace.ofCore.normSq x` etc (defeq to `is_R_or_C.re ⟪x, x⟫`) instead of `‖x‖ ^ 2` +etc to avoid extra rewrites when applying it to an `InnerProductSpace`. -/ +theorem cauchy_schwarz_aux (x y : F) : normSqF (⟪x, y⟫ • x - ⟪x, x⟫ • y) + = normSqF x * (normSqF x * normSqF y - ‖⟪x, y⟫‖ ^ 2) := by + rw [← @ofReal_inj 𝕜, ofReal_normSq_eq_inner_self] + simp only [inner_sub_sub_self, inner_smul_left, inner_smul_right, conj_ofReal, mul_sub, ← + ofReal_normSq_eq_inner_self x, ← ofReal_normSq_eq_inner_self y] + rw [← mul_assoc, mul_conj, RCLike.conj_mul, mul_left_comm, ← inner_conj_symm y, mul_conj] + push_cast + ring + +/-- **Cauchy–Schwarz inequality**. +We need this for the `PreInnerProductSpace.Core` structure to prove the triangle inequality below +when showing the core is a normed group and to take the quotient. + +(This is not intended for general use; see `Analysis.InnerProductSpace.Basic` for a variety of +versions of Cauchy-Schwartz for an inner product space, rather than a `PreInnerProductSpace.Core`). +-/ +theorem inner_mul_inner_self_le (x y : F) : ‖⟪x, y⟫‖ * ‖⟪y, x⟫‖ ≤ re ⟪x, x⟫ * re ⟪y, y⟫ := by + suffices discrim (normSqF x) (2 * ‖⟪x, y⟫_𝕜‖) (normSqF y) ≤ 0 by + rw [norm_inner_symm y x] + rw [discrim, normSq, normSq, sq] at this + linarith + refine discrim_le_zero fun t ↦ ?_ + by_cases hzero : ⟪x, y⟫ = 0 + · simp only [mul_assoc, ← sq, hzero, norm_zero, mul_zero, zero_mul, add_zero, ge_iff_le] + obtain ⟨hx, hy⟩ : (0 ≤ normSqF x ∧ 0 ≤ normSqF y) := ⟨inner_self_nonneg, inner_self_nonneg⟩ + positivity + · have hzero' : ‖⟪x, y⟫‖ ≠ 0 := norm_ne_zero_iff.2 hzero + convert cauchy_schwarz_aux' (𝕜 := 𝕜) (⟪x, y⟫ • x) y (t / ‖⟪x, y⟫‖) using 3 + · field_simp + rw [← sq, normSq, normSq, inner_smul_right, inner_smul_left, ← mul_assoc _ _ ⟪x, x⟫, + mul_conj] + nth_rw 2 [sq] + rw [← ofReal_mul, re_ofReal_mul] + ring + · field_simp + rw [inner_smul_left, mul_comm _ ⟪x, y⟫_𝕜, mul_conj, ← ofReal_pow, ofReal_re] + ring + +/-- (Semi)norm constructed from an `PreInnerProductSpace.Core` structure, defined to be the square +root of the scalar product. -/ +def toNorm : Norm F where norm x := √(re ⟪x, x⟫) + +attribute [local instance] toNorm + +theorem norm_eq_sqrt_inner (x : F) : ‖x‖ = √(re ⟪x, x⟫) := rfl + +theorem inner_self_eq_norm_mul_norm (x : F) : re ⟪x, x⟫ = ‖x‖ * ‖x‖ := by + rw [norm_eq_sqrt_inner, ← sqrt_mul inner_self_nonneg (re ⟪x, x⟫), sqrt_mul_self inner_self_nonneg] + +theorem sqrt_normSq_eq_norm (x : F) : √(normSqF x) = ‖x‖ := rfl + +/-- Cauchy–Schwarz inequality with norm -/ +theorem norm_inner_le_norm (x y : F) : ‖⟪x, y⟫‖ ≤ ‖x‖ * ‖y‖ := + nonneg_le_nonneg_of_sq_le_sq (mul_nonneg (sqrt_nonneg _) (sqrt_nonneg _)) <| + calc + ‖⟪x, y⟫‖ * ‖⟪x, y⟫‖ = ‖⟪x, y⟫‖ * ‖⟪y, x⟫‖ := by rw [norm_inner_symm] + _ ≤ re ⟪x, x⟫ * re ⟪y, y⟫ := inner_mul_inner_self_le x y + _ = ‖x‖ * ‖y‖ * (‖x‖ * ‖y‖) := by simp only [inner_self_eq_norm_mul_norm]; ring + +/-- Seminormed group structure constructed from an `PreInnerProductSpace.Core` structure -/ +def toSeminormedAddCommGroup : SeminormedAddCommGroup F := + AddGroupSeminorm.toSeminormedAddCommGroup + { toFun := fun x => √(re ⟪x, x⟫) + map_zero' := by simp only [sqrt_zero, inner_zero_right, map_zero] + neg' := fun x => by simp only [inner_neg_left, neg_neg, inner_neg_right] + add_le' := fun x y => by + have h₁ : ‖⟪x, y⟫‖ ≤ ‖x‖ * ‖y‖ := norm_inner_le_norm _ _ + have h₂ : re ⟪x, y⟫ ≤ ‖⟪x, y⟫‖ := re_le_norm _ + have h₃ : re ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := h₂.trans h₁ + have h₄ : re ⟪y, x⟫ ≤ ‖x‖ * ‖y‖ := by rwa [← inner_conj_symm, conj_re] + have : ‖x + y‖ * ‖x + y‖ ≤ (‖x‖ + ‖y‖) * (‖x‖ + ‖y‖) := by + simp only [← inner_self_eq_norm_mul_norm, inner_add_add_self, mul_add, mul_comm, map_add] + linarith + exact nonneg_le_nonneg_of_sq_le_sq (add_nonneg (sqrt_nonneg _) (sqrt_nonneg _)) this } + +attribute [local instance] toSeminormedAddCommGroup + +/-- Normed space (which is actually a seminorm) structure constructed from an +`PreInnerProductSpace.Core` structure -/ +def toSeminormedSpace : NormedSpace 𝕜 F where + norm_smul_le r x := by + rw [norm_eq_sqrt_inner, inner_smul_left, inner_smul_right, ← mul_assoc] + rw [RCLike.conj_mul, ← ofReal_pow, re_ofReal_mul, sqrt_mul, ← ofReal_normSq_eq_inner_self, + ofReal_re] + · simp [sqrt_normSq_eq_norm, RCLike.sqrt_normSq_eq_norm] + · positivity + +end PreInnerProductSpace.Core + +section InnerProductSpace.Core + +variable [AddCommGroup F] [Module 𝕜 F] [cd : InnerProductSpace.Core 𝕜 F] + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 F _ x y + +local notation "normSqK" => @RCLike.normSq 𝕜 _ + +local notation "ext_iff" => @RCLike.ext_iff 𝕜 _ + +/-- Inner product defined by the `InnerProductSpace.Core` structure. We can't reuse +`InnerProductSpace.Core.toInner` because it takes `InnerProductSpace.Core` as an explicit +argument. -/ +def toInner' : Inner 𝕜 F := + cd.toInner + +attribute [local instance] toInner' + +local notation "normSqF" => @normSq 𝕜 F _ _ _ _ + +theorem inner_self_eq_zero {x : F} : ⟪x, x⟫ = 0 ↔ x = 0 := + ⟨cd.definite _, inner_self_of_eq_zero⟩ + +theorem normSq_eq_zero {x : F} : normSqF x = 0 ↔ x = 0 := + Iff.trans + (by simp only [normSq, ext_iff, map_zero, inner_self_im, eq_self_iff_true, and_true]) + (@inner_self_eq_zero 𝕜 _ _ _ _ _ x) + +theorem inner_self_ne_zero {x : F} : ⟪x, x⟫ ≠ 0 ↔ x ≠ 0 := + inner_self_eq_zero.not + +attribute [local instance] toNorm + +/-- Normed group structure constructed from an `InnerProductSpace.Core` structure -/ +def toNormedAddCommGroup : NormedAddCommGroup F := + AddGroupNorm.toNormedAddCommGroup + { toFun := fun x => √(re ⟪x, x⟫) + map_zero' := by simp only [sqrt_zero, inner_zero_right, map_zero] + neg' := fun x => by simp only [inner_neg_left, neg_neg, inner_neg_right] + add_le' := fun x y => by + have h₁ : ‖⟪x, y⟫‖ ≤ ‖x‖ * ‖y‖ := norm_inner_le_norm _ _ + have h₂ : re ⟪x, y⟫ ≤ ‖⟪x, y⟫‖ := re_le_norm _ + have h₃ : re ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := h₂.trans h₁ + have h₄ : re ⟪y, x⟫ ≤ ‖x‖ * ‖y‖ := by rwa [← inner_conj_symm, conj_re] + have : ‖x + y‖ * ‖x + y‖ ≤ (‖x‖ + ‖y‖) * (‖x‖ + ‖y‖) := by + simp only [← inner_self_eq_norm_mul_norm, inner_add_add_self, mul_add, mul_comm, map_add] + linarith + exact nonneg_le_nonneg_of_sq_le_sq (add_nonneg (sqrt_nonneg _) (sqrt_nonneg _)) this + eq_zero_of_map_eq_zero' := fun _ hx => + normSq_eq_zero.1 <| (sqrt_eq_zero inner_self_nonneg).1 hx } + +attribute [local instance] toNormedAddCommGroup + +/-- Normed space structure constructed from an `InnerProductSpace.Core` structure -/ +def toNormedSpace : NormedSpace 𝕜 F where + norm_smul_le r x := by + rw [norm_eq_sqrt_inner, inner_smul_left, inner_smul_right, ← mul_assoc] + rw [RCLike.conj_mul, ← ofReal_pow, re_ofReal_mul, sqrt_mul, ← ofReal_normSq_eq_inner_self, + ofReal_re] + · simp [sqrt_normSq_eq_norm, RCLike.sqrt_normSq_eq_norm] + · positivity + +end InnerProductSpace.Core + +end InnerProductSpace.Core + +section + +attribute [local instance] InnerProductSpace.Core.toNormedAddCommGroup + +/-- Given an `InnerProductSpace.Core` structure on a space, one can use it to turn +the space into an inner product space. The `NormedAddCommGroup` structure is expected +to already be defined with `InnerProductSpace.ofCore.toNormedAddCommGroup`. -/ +def InnerProductSpace.ofCore [AddCommGroup F] [Module 𝕜 F] (cd : InnerProductSpace.Core 𝕜 F) : + InnerProductSpace 𝕜 F := + letI : NormedSpace 𝕜 F := @InnerProductSpace.Core.toNormedSpace 𝕜 F _ _ _ cd + { cd with + norm_sq_eq_inner := fun x => by + have h₁ : ‖x‖ ^ 2 = √(re (cd.inner x x)) ^ 2 := rfl + have h₂ : 0 ≤ re (cd.inner x x) := InnerProductSpace.Core.inner_self_nonneg + simp [h₁, sq_sqrt, h₂] } + +end + +end diff --git a/Mathlib/Analysis/InnerProductSpace/LinearMap.lean b/Mathlib/Analysis/InnerProductSpace/LinearMap.lean new file mode 100644 index 0000000000000..c04c6c4fd1d7f --- /dev/null +++ b/Mathlib/Analysis/InnerProductSpace/LinearMap.lean @@ -0,0 +1,308 @@ +/- +Copyright (c) 2019 Zhouhang Zhou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Zhouhang Zhou, Sébastien Gouëzel, Frédéric Dupuis +-/ + +import Mathlib.Analysis.InnerProductSpace.Basic + +/-! +# Linear maps on inner product spaces + +This file studies linear maps on inner product spaces. + +## Main results + +- We define `innerSL` as the inner product bundled as a continuous sesquilinear map, and + `innerₛₗ` for the non-continuous version. +- We prove a general polarization identity for linear maps (`inner_map_polarization`) +- We show that a linear map preserving the inner product is an isometry + (`LinearMap.isometryOfInner`) and conversely an isometry preserves the inner product + (`LinearIsometry.inner_map_map`). + +## Tags + +inner product space, Hilbert space, norm + +-/ + +noncomputable section + +open RCLike Real Filter Topology ComplexConjugate Finsupp + +open LinearMap (BilinForm) + +variable {𝕜 E F : Type*} [RCLike 𝕜] + +section Norm_Seminormed + +open scoped InnerProductSpace + +variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] +variable [SeminormedAddCommGroup F] [InnerProductSpace ℝ F] + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + +section Complex_Seminormed + +variable {V : Type*} [SeminormedAddCommGroup V] [InnerProductSpace ℂ V] + +/-- A complex polarization identity, with a linear map. -/ +theorem inner_map_polarization (T : V →ₗ[ℂ] V) (x y : V) : + ⟪T y, x⟫_ℂ = + (⟪T (x + y), x + y⟫_ℂ - ⟪T (x - y), x - y⟫_ℂ + + Complex.I * ⟪T (x + Complex.I • y), x + Complex.I • y⟫_ℂ - + Complex.I * ⟪T (x - Complex.I • y), x - Complex.I • y⟫_ℂ) / + 4 := by + simp only [map_add, map_sub, inner_add_left, inner_add_right, LinearMap.map_smul, inner_smul_left, + inner_smul_right, Complex.conj_I, ← pow_two, Complex.I_sq, inner_sub_left, inner_sub_right, + mul_add, ← mul_assoc, mul_neg, neg_neg, sub_neg_eq_add, one_mul, neg_one_mul, mul_sub, sub_sub] + ring + +theorem inner_map_polarization' (T : V →ₗ[ℂ] V) (x y : V) : + ⟪T x, y⟫_ℂ = + (⟪T (x + y), x + y⟫_ℂ - ⟪T (x - y), x - y⟫_ℂ - + Complex.I * ⟪T (x + Complex.I • y), x + Complex.I • y⟫_ℂ + + Complex.I * ⟪T (x - Complex.I • y), x - Complex.I • y⟫_ℂ) / + 4 := by + simp only [map_add, map_sub, inner_add_left, inner_add_right, LinearMap.map_smul, inner_smul_left, + inner_smul_right, Complex.conj_I, ← pow_two, Complex.I_sq, inner_sub_left, inner_sub_right, + mul_add, ← mul_assoc, mul_neg, neg_neg, sub_neg_eq_add, one_mul, neg_one_mul, mul_sub, sub_sub] + ring + +end Complex_Seminormed + +section Complex + +variable {V : Type*} [NormedAddCommGroup V] [InnerProductSpace ℂ V] + +/-- A linear map `T` is zero, if and only if the identity `⟪T x, x⟫_ℂ = 0` holds for all `x`. +-/ +theorem inner_map_self_eq_zero (T : V →ₗ[ℂ] V) : (∀ x : V, ⟪T x, x⟫_ℂ = 0) ↔ T = 0 := by + constructor + · intro hT + ext x + rw [LinearMap.zero_apply, ← @inner_self_eq_zero ℂ V, inner_map_polarization] + simp only [hT] + norm_num + · rintro rfl x + simp only [LinearMap.zero_apply, inner_zero_left] + +/-- +Two linear maps `S` and `T` are equal, if and only if the identity `⟪S x, x⟫_ℂ = ⟪T x, x⟫_ℂ` holds +for all `x`. +-/ +theorem ext_inner_map (S T : V →ₗ[ℂ] V) : (∀ x : V, ⟪S x, x⟫_ℂ = ⟪T x, x⟫_ℂ) ↔ S = T := by + rw [← sub_eq_zero, ← inner_map_self_eq_zero] + refine forall_congr' fun x => ?_ + rw [LinearMap.sub_apply, inner_sub_left, sub_eq_zero] + +end Complex + +section + +variable {ι : Type*} {ι' : Type*} {ι'' : Type*} +variable {E' : Type*} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] +variable {E'' : Type*} [SeminormedAddCommGroup E''] [InnerProductSpace 𝕜 E''] + +/-- A linear isometry preserves the inner product. -/ +@[simp] +theorem LinearIsometry.inner_map_map (f : E →ₗᵢ[𝕜] E') (x y : E) : ⟪f x, f y⟫ = ⟪x, y⟫ := by + simp [inner_eq_sum_norm_sq_div_four, ← f.norm_map] + +/-- A linear isometric equivalence preserves the inner product. -/ +@[simp] +theorem LinearIsometryEquiv.inner_map_map (f : E ≃ₗᵢ[𝕜] E') (x y : E) : ⟪f x, f y⟫ = ⟪x, y⟫ := + f.toLinearIsometry.inner_map_map x y + +/-- The adjoint of a linear isometric equivalence is its inverse. -/ +theorem LinearIsometryEquiv.inner_map_eq_flip (f : E ≃ₗᵢ[𝕜] E') (x : E) (y : E') : + ⟪f x, y⟫_𝕜 = ⟪x, f.symm y⟫_𝕜 := by + conv_lhs => rw [← f.apply_symm_apply y, f.inner_map_map] + +/-- A linear map that preserves the inner product is a linear isometry. -/ +def LinearMap.isometryOfInner (f : E →ₗ[𝕜] E') (h : ∀ x y, ⟪f x, f y⟫ = ⟪x, y⟫) : E →ₗᵢ[𝕜] E' := + ⟨f, fun x => by simp only [@norm_eq_sqrt_inner 𝕜, h]⟩ + +@[simp] +theorem LinearMap.coe_isometryOfInner (f : E →ₗ[𝕜] E') (h) : ⇑(f.isometryOfInner h) = f := + rfl + +@[simp] +theorem LinearMap.isometryOfInner_toLinearMap (f : E →ₗ[𝕜] E') (h) : + (f.isometryOfInner h).toLinearMap = f := + rfl + +/-- A linear equivalence that preserves the inner product is a linear isometric equivalence. -/ +def LinearEquiv.isometryOfInner (f : E ≃ₗ[𝕜] E') (h : ∀ x y, ⟪f x, f y⟫ = ⟪x, y⟫) : E ≃ₗᵢ[𝕜] E' := + ⟨f, ((f : E →ₗ[𝕜] E').isometryOfInner h).norm_map⟩ + +@[simp] +theorem LinearEquiv.coe_isometryOfInner (f : E ≃ₗ[𝕜] E') (h) : ⇑(f.isometryOfInner h) = f := + rfl + +@[simp] +theorem LinearEquiv.isometryOfInner_toLinearEquiv (f : E ≃ₗ[𝕜] E') (h) : + (f.isometryOfInner h).toLinearEquiv = f := + rfl + +/-- A linear map is an isometry if and it preserves the inner product. -/ +theorem LinearMap.norm_map_iff_inner_map_map {F : Type*} [FunLike F E E'] [LinearMapClass F 𝕜 E E'] + (f : F) : (∀ x, ‖f x‖ = ‖x‖) ↔ (∀ x y, ⟪f x, f y⟫_𝕜 = ⟪x, y⟫_𝕜) := + ⟨({ toLinearMap := LinearMapClass.linearMap f, norm_map' := · : E →ₗᵢ[𝕜] E' }.inner_map_map), + (LinearMapClass.linearMap f |>.isometryOfInner · |>.norm_map)⟩ + +end + +variable (𝕜) + +/-- The inner product as a sesquilinear map. -/ +def innerₛₗ : E →ₗ⋆[𝕜] E →ₗ[𝕜] 𝕜 := + LinearMap.mk₂'ₛₗ _ _ (fun v w => ⟪v, w⟫) inner_add_left (fun _ _ _ => inner_smul_left _ _ _) + inner_add_right fun _ _ _ => inner_smul_right _ _ _ + +@[simp] +theorem innerₛₗ_apply_coe (v : E) : ⇑(innerₛₗ 𝕜 v) = fun w => ⟪v, w⟫ := + rfl + +@[simp] +theorem innerₛₗ_apply (v w : E) : innerₛₗ 𝕜 v w = ⟪v, w⟫ := + rfl + +variable (F) +/-- The inner product as a bilinear map in the real case. -/ +def innerₗ : F →ₗ[ℝ] F →ₗ[ℝ] ℝ := innerₛₗ ℝ + +@[simp] lemma flip_innerₗ : (innerₗ F).flip = innerₗ F := by + ext v w + exact real_inner_comm v w + +variable {F} + +@[simp] lemma innerₗ_apply (v w : F) : innerₗ F v w = ⟪v, w⟫_ℝ := rfl + +/-- The inner product as a continuous sesquilinear map. Note that `toDualMap` (resp. `toDual`) +in `InnerProductSpace.Dual` is a version of this given as a linear isometry (resp. linear +isometric equivalence). -/ +def innerSL : E →L⋆[𝕜] E →L[𝕜] 𝕜 := + LinearMap.mkContinuous₂ (innerₛₗ 𝕜) 1 fun x y => by + simp only [norm_inner_le_norm, one_mul, innerₛₗ_apply] + +@[simp] +theorem innerSL_apply_coe (v : E) : ⇑(innerSL 𝕜 v) = fun w => ⟪v, w⟫ := + rfl + +@[simp] +theorem innerSL_apply (v w : E) : innerSL 𝕜 v w = ⟪v, w⟫ := + rfl + +/-- The inner product as a continuous sesquilinear map, with the two arguments flipped. -/ +def innerSLFlip : E →L[𝕜] E →L⋆[𝕜] 𝕜 := + @ContinuousLinearMap.flipₗᵢ' 𝕜 𝕜 𝕜 E E 𝕜 _ _ _ _ _ _ _ _ _ (RingHom.id 𝕜) (starRingEnd 𝕜) _ _ + (innerSL 𝕜) + +@[simp] +theorem innerSLFlip_apply (x y : E) : innerSLFlip 𝕜 x y = ⟪y, x⟫ := + rfl + +variable (F) in +@[simp] lemma innerSL_real_flip : (innerSL ℝ (E := F)).flip = innerSL ℝ := by + ext v w + exact real_inner_comm _ _ + +variable {𝕜} + +namespace ContinuousLinearMap + +variable {E' : Type*} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] + +-- Note: odd and expensive build behavior is explicitly turned off using `noncomputable` +/-- Given `f : E →L[𝕜] E'`, construct the continuous sesquilinear form `fun x y ↦ ⟪x, A y⟫`, given +as a continuous linear map. -/ +noncomputable def toSesqForm : (E →L[𝕜] E') →L[𝕜] E' →L⋆[𝕜] E →L[𝕜] 𝕜 := + (ContinuousLinearMap.flipₗᵢ' E E' 𝕜 (starRingEnd 𝕜) (RingHom.id 𝕜)).toContinuousLinearEquiv ∘L + ContinuousLinearMap.compSL E E' (E' →L⋆[𝕜] 𝕜) (RingHom.id 𝕜) (RingHom.id 𝕜) (innerSLFlip 𝕜) + +@[simp] +theorem toSesqForm_apply_coe (f : E →L[𝕜] E') (x : E') : toSesqForm f x = (innerSL 𝕜 x).comp f := + rfl + +theorem toSesqForm_apply_norm_le {f : E →L[𝕜] E'} {v : E'} : ‖toSesqForm f v‖ ≤ ‖f‖ * ‖v‖ := by + refine opNorm_le_bound _ (by positivity) fun x ↦ ?_ + have h₁ : ‖f x‖ ≤ ‖f‖ * ‖x‖ := le_opNorm _ _ + have h₂ := @norm_inner_le_norm 𝕜 E' _ _ _ v (f x) + calc + ‖⟪v, f x⟫‖ ≤ ‖v‖ * ‖f x‖ := h₂ + _ ≤ ‖v‖ * (‖f‖ * ‖x‖) := mul_le_mul_of_nonneg_left h₁ (norm_nonneg v) + _ = ‖f‖ * ‖v‖ * ‖x‖ := by ring + +end ContinuousLinearMap + +variable (𝕜) + +/-- `innerSL` is an isometry. Note that the associated `LinearIsometry` is defined in +`InnerProductSpace.Dual` as `toDualMap`. -/ +@[simp] +theorem innerSL_apply_norm (x : E) : ‖innerSL 𝕜 x‖ = ‖x‖ := by + refine + le_antisymm ((innerSL 𝕜 x).opNorm_le_bound (norm_nonneg _) fun y => norm_inner_le_norm _ _) ?_ + rcases (norm_nonneg x).eq_or_gt with (h | h) + · simp [h] + · refine (mul_le_mul_right h).mp ?_ + calc + ‖x‖ * ‖x‖ = ‖(⟪x, x⟫ : 𝕜)‖ := by + rw [← sq, inner_self_eq_norm_sq_to_K, norm_pow, norm_ofReal, abs_norm] + _ ≤ ‖innerSL 𝕜 x‖ * ‖x‖ := (innerSL 𝕜 x).le_opNorm _ + +lemma norm_innerSL_le : ‖innerSL 𝕜 (E := E)‖ ≤ 1 := + ContinuousLinearMap.opNorm_le_bound _ zero_le_one (by simp) + +end Norm_Seminormed + +section RCLikeToReal + +open scoped InnerProductSpace + +variable {G : Type*} + +/-- The inner product on an inner product space of dimension 2 can be evaluated in terms +of a complex-number representation of the space. -/ +theorem inner_map_complex [SeminormedAddCommGroup G] [InnerProductSpace ℝ G] (f : G ≃ₗᵢ[ℝ] ℂ) + (x y : G) : ⟪x, y⟫_ℝ = (conj (f x) * f y).re := by rw [← Complex.inner, f.inner_map_map] + +end RCLikeToReal + +section ReApplyInnerSelf + +variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + +/-- Extract a real bilinear form from an operator `T`, +by taking the pairing `fun x ↦ re ⟪T x, x⟫`. -/ +def ContinuousLinearMap.reApplyInnerSelf (T : E →L[𝕜] E) (x : E) : ℝ := + re ⟪T x, x⟫ + +theorem ContinuousLinearMap.reApplyInnerSelf_apply (T : E →L[𝕜] E) (x : E) : + T.reApplyInnerSelf x = re ⟪T x, x⟫ := + rfl + +end ReApplyInnerSelf + +section ReApplyInnerSelf_Seminormed + +variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + +theorem ContinuousLinearMap.reApplyInnerSelf_continuous (T : E →L[𝕜] E) : + Continuous T.reApplyInnerSelf := + reCLM.continuous.comp <| T.continuous.inner continuous_id + +theorem ContinuousLinearMap.reApplyInnerSelf_smul (T : E →L[𝕜] E) (x : E) {c : 𝕜} : + T.reApplyInnerSelf (c • x) = ‖c‖ ^ 2 * T.reApplyInnerSelf x := by + simp only [ContinuousLinearMap.map_smul, ContinuousLinearMap.reApplyInnerSelf_apply, + inner_smul_left, inner_smul_right, ← mul_assoc, mul_conj, ← ofReal_pow, ← smul_re, + Algebra.smul_def (‖c‖ ^ 2) ⟪T x, x⟫, algebraMap_eq_ofReal] + +end ReApplyInnerSelf_Seminormed diff --git a/Mathlib/Analysis/InnerProductSpace/Orthogonal.lean b/Mathlib/Analysis/InnerProductSpace/Orthogonal.lean index 7b857c4ae5eba..70a253a0c3e16 100644 --- a/Mathlib/Analysis/InnerProductSpace/Orthogonal.lean +++ b/Mathlib/Analysis/InnerProductSpace/Orthogonal.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Zhouhang Zhou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Zhouhang Zhou, Sébastien Gouëzel, Frédéric Dupuis -/ -import Mathlib.Analysis.InnerProductSpace.Basic +import Mathlib.Analysis.InnerProductSpace.Subspace import Mathlib.LinearAlgebra.SesquilinearForm /-! diff --git a/Mathlib/Analysis/InnerProductSpace/Orthonormal.lean b/Mathlib/Analysis/InnerProductSpace/Orthonormal.lean new file mode 100644 index 0000000000000..2013121807f71 --- /dev/null +++ b/Mathlib/Analysis/InnerProductSpace/Orthonormal.lean @@ -0,0 +1,444 @@ +/- +Copyright (c) 2019 Zhouhang Zhou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Zhouhang Zhou, Sébastien Gouëzel, Frédéric Dupuis +-/ + +import Mathlib.Analysis.InnerProductSpace.LinearMap + +/-! +# Orthonormal sets + +This file defines orthonormal sets in inner product spaces. + +## Main results + +- We define `Orthonormal`, a predicate on a function `v : ι → E`, and prove the existence of a + maximal orthonormal set, `exists_maximal_orthonormal`. +- Bessel's inequality, `Orthonormal.tsum_inner_products_le`, states that given an orthonormal set + `v` and a vector `x`, the sum of the norm-squares of the inner products `⟪v i, x⟫` is no more + than the norm-square of `x`. + +For the existence of orthonormal bases, Hilbert bases, etc., see the file +`Analysis.InnerProductSpace.projection`. +-/ + +noncomputable section + +open RCLike Real Filter + +open Topology ComplexConjugate Finsupp + +open LinearMap (BilinForm) + +variable {𝕜 E F : Type*} [RCLike 𝕜] + +section OrthonormalSets_Seminormed + +variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] +variable [SeminormedAddCommGroup F] [InnerProductSpace ℝ F] + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + +variable {ι : Type*} (𝕜) + +/-- An orthonormal set of vectors in an `InnerProductSpace` -/ +def Orthonormal (v : ι → E) : Prop := + (∀ i, ‖v i‖ = 1) ∧ Pairwise fun i j => ⟪v i, v j⟫ = 0 + +variable {𝕜} + +/-- `if ... then ... else` characterization of an indexed set of vectors being orthonormal. (Inner +product equals Kronecker delta.) -/ +theorem orthonormal_iff_ite [DecidableEq ι] {v : ι → E} : + Orthonormal 𝕜 v ↔ ∀ i j, ⟪v i, v j⟫ = if i = j then (1 : 𝕜) else (0 : 𝕜) := by + constructor + · intro hv i j + split_ifs with h + · simp [h, inner_self_eq_norm_sq_to_K, hv.1] + · exact hv.2 h + · intro h + constructor + · intro i + have h' : ‖v i‖ ^ 2 = 1 ^ 2 := by simp [@norm_sq_eq_inner 𝕜, h i i] + have h₁ : 0 ≤ ‖v i‖ := norm_nonneg _ + have h₂ : (0 : ℝ) ≤ 1 := zero_le_one + rwa [sq_eq_sq₀ h₁ h₂] at h' + · intro i j hij + simpa [hij] using h i j + +/-- `if ... then ... else` characterization of a set of vectors being orthonormal. (Inner product +equals Kronecker delta.) -/ +theorem orthonormal_subtype_iff_ite [DecidableEq E] {s : Set E} : + Orthonormal 𝕜 (Subtype.val : s → E) ↔ ∀ v ∈ s, ∀ w ∈ s, ⟪v, w⟫ = if v = w then 1 else 0 := by + rw [orthonormal_iff_ite] + constructor + · intro h v hv w hw + convert h ⟨v, hv⟩ ⟨w, hw⟩ using 1 + simp + · rintro h ⟨v, hv⟩ ⟨w, hw⟩ + convert h v hv w hw using 1 + simp + +/-- The inner product of a linear combination of a set of orthonormal vectors with one of those +vectors picks out the coefficient of that vector. -/ +theorem Orthonormal.inner_right_finsupp {v : ι → E} (hv : Orthonormal 𝕜 v) (l : ι →₀ 𝕜) (i : ι) : + ⟪v i, linearCombination 𝕜 v l⟫ = l i := by + classical + simpa [linearCombination_apply, Finsupp.inner_sum, orthonormal_iff_ite.mp hv] using Eq.symm + +/-- The inner product of a linear combination of a set of orthonormal vectors with one of those +vectors picks out the coefficient of that vector. -/ +theorem Orthonormal.inner_right_sum {v : ι → E} (hv : Orthonormal 𝕜 v) (l : ι → 𝕜) {s : Finset ι} + {i : ι} (hi : i ∈ s) : ⟪v i, ∑ i ∈ s, l i • v i⟫ = l i := by + classical + simp [inner_sum, inner_smul_right, orthonormal_iff_ite.mp hv, hi] + +/-- The inner product of a linear combination of a set of orthonormal vectors with one of those +vectors picks out the coefficient of that vector. -/ +theorem Orthonormal.inner_right_fintype [Fintype ι] {v : ι → E} (hv : Orthonormal 𝕜 v) (l : ι → 𝕜) + (i : ι) : ⟪v i, ∑ i : ι, l i • v i⟫ = l i := + hv.inner_right_sum l (Finset.mem_univ _) + +/-- The inner product of a linear combination of a set of orthonormal vectors with one of those +vectors picks out the coefficient of that vector. -/ +theorem Orthonormal.inner_left_finsupp {v : ι → E} (hv : Orthonormal 𝕜 v) (l : ι →₀ 𝕜) (i : ι) : + ⟪linearCombination 𝕜 v l, v i⟫ = conj (l i) := by rw [← inner_conj_symm, hv.inner_right_finsupp] + +/-- The inner product of a linear combination of a set of orthonormal vectors with one of those +vectors picks out the coefficient of that vector. -/ +theorem Orthonormal.inner_left_sum {v : ι → E} (hv : Orthonormal 𝕜 v) (l : ι → 𝕜) {s : Finset ι} + {i : ι} (hi : i ∈ s) : ⟪∑ i ∈ s, l i • v i, v i⟫ = conj (l i) := by + classical + simp only [sum_inner, inner_smul_left, orthonormal_iff_ite.mp hv, hi, mul_boole, + Finset.sum_ite_eq', if_true] + +/-- The inner product of a linear combination of a set of orthonormal vectors with one of those +vectors picks out the coefficient of that vector. -/ +theorem Orthonormal.inner_left_fintype [Fintype ι] {v : ι → E} (hv : Orthonormal 𝕜 v) (l : ι → 𝕜) + (i : ι) : ⟪∑ i : ι, l i • v i, v i⟫ = conj (l i) := + hv.inner_left_sum l (Finset.mem_univ _) + +/-- The inner product of two linear combinations of a set of orthonormal vectors, expressed as +a sum over the first `Finsupp`. -/ +theorem Orthonormal.inner_finsupp_eq_sum_left {v : ι → E} (hv : Orthonormal 𝕜 v) (l₁ l₂ : ι →₀ 𝕜) : + ⟪linearCombination 𝕜 v l₁, linearCombination 𝕜 v l₂⟫ = l₁.sum fun i y => conj y * l₂ i := by + simp only [l₁.linearCombination_apply _, Finsupp.sum_inner, hv.inner_right_finsupp, smul_eq_mul] + +/-- The inner product of two linear combinations of a set of orthonormal vectors, expressed as +a sum over the second `Finsupp`. -/ +theorem Orthonormal.inner_finsupp_eq_sum_right {v : ι → E} (hv : Orthonormal 𝕜 v) (l₁ l₂ : ι →₀ 𝕜) : + ⟪linearCombination 𝕜 v l₁, linearCombination 𝕜 v l₂⟫ = l₂.sum fun i y => conj (l₁ i) * y := by + simp only [l₂.linearCombination_apply _, Finsupp.inner_sum, hv.inner_left_finsupp, mul_comm, + smul_eq_mul] + +/-- The inner product of two linear combinations of a set of orthonormal vectors, expressed as +a sum. -/ +protected theorem Orthonormal.inner_sum {v : ι → E} (hv : Orthonormal 𝕜 v) (l₁ l₂ : ι → 𝕜) + (s : Finset ι) : ⟪∑ i ∈ s, l₁ i • v i, ∑ i ∈ s, l₂ i • v i⟫ = ∑ i ∈ s, conj (l₁ i) * l₂ i := by + simp_rw [sum_inner, inner_smul_left] + refine Finset.sum_congr rfl fun i hi => ?_ + rw [hv.inner_right_sum l₂ hi] + +/-- +The double sum of weighted inner products of pairs of vectors from an orthonormal sequence is the +sum of the weights. +-/ +theorem Orthonormal.inner_left_right_finset {s : Finset ι} {v : ι → E} (hv : Orthonormal 𝕜 v) + {a : ι → ι → 𝕜} : (∑ i ∈ s, ∑ j ∈ s, a i j • ⟪v j, v i⟫) = ∑ k ∈ s, a k k := by + classical + simp [orthonormal_iff_ite.mp hv, Finset.sum_ite_of_true] + +/-- An orthonormal set is linearly independent. -/ +theorem Orthonormal.linearIndependent {v : ι → E} (hv : Orthonormal 𝕜 v) : + LinearIndependent 𝕜 v := by + rw [linearIndependent_iff] + intro l hl + ext i + have key : ⟪v i, Finsupp.linearCombination 𝕜 v l⟫ = ⟪v i, 0⟫ := by rw [hl] + simpa only [hv.inner_right_finsupp, inner_zero_right] using key + +/-- A subfamily of an orthonormal family (i.e., a composition with an injective map) is an +orthonormal family. -/ +theorem Orthonormal.comp {ι' : Type*} {v : ι → E} (hv : Orthonormal 𝕜 v) (f : ι' → ι) + (hf : Function.Injective f) : Orthonormal 𝕜 (v ∘ f) := by + classical + rw [orthonormal_iff_ite] at hv ⊢ + intro i j + convert hv (f i) (f j) using 1 + simp [hf.eq_iff] + +/-- An injective family `v : ι → E` is orthonormal if and only if `Subtype.val : (range v) → E` is +orthonormal. -/ +theorem orthonormal_subtype_range {v : ι → E} (hv : Function.Injective v) : + Orthonormal 𝕜 (Subtype.val : Set.range v → E) ↔ Orthonormal 𝕜 v := by + let f : ι ≃ Set.range v := Equiv.ofInjective v hv + refine ⟨fun h => h.comp f f.injective, fun h => ?_⟩ + rw [← Equiv.self_comp_ofInjective_symm hv] + exact h.comp f.symm f.symm.injective + +/-- If `v : ι → E` is an orthonormal family, then `Subtype.val : (range v) → E` is an orthonormal +family. -/ +theorem Orthonormal.toSubtypeRange {v : ι → E} (hv : Orthonormal 𝕜 v) : + Orthonormal 𝕜 (Subtype.val : Set.range v → E) := + (orthonormal_subtype_range hv.linearIndependent.injective).2 hv + +/-- A linear combination of some subset of an orthonormal set is orthogonal to other members of the +set. -/ +theorem Orthonormal.inner_finsupp_eq_zero {v : ι → E} (hv : Orthonormal 𝕜 v) {s : Set ι} {i : ι} + (hi : i ∉ s) {l : ι →₀ 𝕜} (hl : l ∈ Finsupp.supported 𝕜 𝕜 s) : + ⟪Finsupp.linearCombination 𝕜 v l, v i⟫ = 0 := by + rw [Finsupp.mem_supported'] at hl + simp only [hv.inner_left_finsupp, hl i hi, map_zero] + +/-- Given an orthonormal family, a second family of vectors is orthonormal if every vector equals +the corresponding vector in the original family or its negation. -/ +theorem Orthonormal.orthonormal_of_forall_eq_or_eq_neg {v w : ι → E} (hv : Orthonormal 𝕜 v) + (hw : ∀ i, w i = v i ∨ w i = -v i) : Orthonormal 𝕜 w := by + classical + rw [orthonormal_iff_ite] at * + intro i j + cases' hw i with hi hi <;> cases' hw j with hj hj <;> + replace hv := hv i j <;> split_ifs at hv ⊢ with h <;> + simpa only [hi, hj, h, inner_neg_right, inner_neg_left, neg_neg, eq_self_iff_true, + neg_eq_zero] using hv + +/- The material that follows, culminating in the existence of a maximal orthonormal subset, is +adapted from the corresponding development of the theory of linearly independents sets. See +`exists_linearIndependent` in particular. -/ +variable (𝕜 E) + +theorem orthonormal_empty : Orthonormal 𝕜 (fun x => x : (∅ : Set E) → E) := by + classical + simp [orthonormal_subtype_iff_ite] + +variable {𝕜 E} + +theorem orthonormal_iUnion_of_directed {η : Type*} {s : η → Set E} (hs : Directed (· ⊆ ·) s) + (h : ∀ i, Orthonormal 𝕜 (fun x => x : s i → E)) : + Orthonormal 𝕜 (fun x => x : (⋃ i, s i) → E) := by + classical + rw [orthonormal_subtype_iff_ite] + rintro x ⟨_, ⟨i, rfl⟩, hxi⟩ y ⟨_, ⟨j, rfl⟩, hyj⟩ + obtain ⟨k, hik, hjk⟩ := hs i j + have h_orth : Orthonormal 𝕜 (fun x => x : s k → E) := h k + rw [orthonormal_subtype_iff_ite] at h_orth + exact h_orth x (hik hxi) y (hjk hyj) + +theorem orthonormal_sUnion_of_directed {s : Set (Set E)} (hs : DirectedOn (· ⊆ ·) s) + (h : ∀ a ∈ s, Orthonormal 𝕜 (fun x => ((x : a) : E))) : + Orthonormal 𝕜 (fun x => x : ⋃₀ s → E) := by + rw [Set.sUnion_eq_iUnion]; exact orthonormal_iUnion_of_directed hs.directed_val (by simpa using h) + +/-- Given an orthonormal set `v` of vectors in `E`, there exists a maximal orthonormal set +containing it. -/ +theorem exists_maximal_orthonormal {s : Set E} (hs : Orthonormal 𝕜 (Subtype.val : s → E)) : + ∃ w ⊇ s, Orthonormal 𝕜 (Subtype.val : w → E) ∧ + ∀ u ⊇ w, Orthonormal 𝕜 (Subtype.val : u → E) → u = w := by + have := zorn_subset_nonempty { b | Orthonormal 𝕜 (Subtype.val : b → E) } ?_ _ hs + · obtain ⟨b, hb⟩ := this + exact ⟨b, hb.1, hb.2.1, fun u hus hu => hb.2.eq_of_ge hu hus ⟩ + · refine fun c hc cc _c0 => ⟨⋃₀ c, ?_, ?_⟩ + · exact orthonormal_sUnion_of_directed cc.directedOn fun x xc => hc xc + · exact fun _ => Set.subset_sUnion_of_mem + +open Module + +/-- A family of orthonormal vectors with the correct cardinality forms a basis. -/ +def basisOfOrthonormalOfCardEqFinrank [Fintype ι] [Nonempty ι] {v : ι → E} (hv : Orthonormal 𝕜 v) + (card_eq : Fintype.card ι = finrank 𝕜 E) : Basis ι 𝕜 E := + basisOfLinearIndependentOfCardEqFinrank hv.linearIndependent card_eq + +@[simp] +theorem coe_basisOfOrthonormalOfCardEqFinrank [Fintype ι] [Nonempty ι] {v : ι → E} + (hv : Orthonormal 𝕜 v) (card_eq : Fintype.card ι = finrank 𝕜 E) : + (basisOfOrthonormalOfCardEqFinrank hv card_eq : ι → E) = v := + coe_basisOfLinearIndependentOfCardEqFinrank _ _ + +theorem Orthonormal.ne_zero {v : ι → E} (hv : Orthonormal 𝕜 v) (i : ι) : v i ≠ 0 := by + refine ne_of_apply_ne norm ?_ + rw [hv.1 i, norm_zero] + norm_num + +end OrthonormalSets_Seminormed + +section Norm_Seminormed + +open scoped InnerProductSpace + +variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] +variable [SeminormedAddCommGroup F] [InnerProductSpace ℝ F] + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + +section + +variable {ι : Type*} {ι' : Type*} {ι'' : Type*} +variable {E' : Type*} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] +variable {E'' : Type*} [SeminormedAddCommGroup E''] [InnerProductSpace 𝕜 E''] + +/-- A linear isometry preserves the property of being orthonormal. -/ +theorem LinearIsometry.orthonormal_comp_iff {v : ι → E} (f : E →ₗᵢ[𝕜] E') : + Orthonormal 𝕜 (f ∘ v) ↔ Orthonormal 𝕜 v := by + classical simp_rw [orthonormal_iff_ite, Function.comp_apply, LinearIsometry.inner_map_map] + +/-- A linear isometry preserves the property of being orthonormal. -/ +theorem Orthonormal.comp_linearIsometry {v : ι → E} (hv : Orthonormal 𝕜 v) (f : E →ₗᵢ[𝕜] E') : + Orthonormal 𝕜 (f ∘ v) := by rwa [f.orthonormal_comp_iff] + +/-- A linear isometric equivalence preserves the property of being orthonormal. -/ +theorem Orthonormal.comp_linearIsometryEquiv {v : ι → E} (hv : Orthonormal 𝕜 v) (f : E ≃ₗᵢ[𝕜] E') : + Orthonormal 𝕜 (f ∘ v) := + hv.comp_linearIsometry f.toLinearIsometry + +/-- A linear isometric equivalence, applied with `Basis.map`, preserves the property of being +orthonormal. -/ +theorem Orthonormal.mapLinearIsometryEquiv {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) + (f : E ≃ₗᵢ[𝕜] E') : Orthonormal 𝕜 (v.map f.toLinearEquiv) := + hv.comp_linearIsometryEquiv f + +/-- A linear map that sends an orthonormal basis to orthonormal vectors is a linear isometry. -/ +def LinearMap.isometryOfOrthonormal (f : E →ₗ[𝕜] E') {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) + (hf : Orthonormal 𝕜 (f ∘ v)) : E →ₗᵢ[𝕜] E' := + f.isometryOfInner fun x y => by + classical rw [← v.linearCombination_repr x, ← v.linearCombination_repr y, + Finsupp.apply_linearCombination, Finsupp.apply_linearCombination, + hv.inner_finsupp_eq_sum_left, hf.inner_finsupp_eq_sum_left] + +@[simp] +theorem LinearMap.coe_isometryOfOrthonormal (f : E →ₗ[𝕜] E') {v : Basis ι 𝕜 E} + (hv : Orthonormal 𝕜 v) (hf : Orthonormal 𝕜 (f ∘ v)) : ⇑(f.isometryOfOrthonormal hv hf) = f := + rfl + +@[simp] +theorem LinearMap.isometryOfOrthonormal_toLinearMap (f : E →ₗ[𝕜] E') {v : Basis ι 𝕜 E} + (hv : Orthonormal 𝕜 v) (hf : Orthonormal 𝕜 (f ∘ v)) : + (f.isometryOfOrthonormal hv hf).toLinearMap = f := + rfl + +/-- A linear equivalence that sends an orthonormal basis to orthonormal vectors is a linear +isometric equivalence. -/ +def LinearEquiv.isometryOfOrthonormal (f : E ≃ₗ[𝕜] E') {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) + (hf : Orthonormal 𝕜 (f ∘ v)) : E ≃ₗᵢ[𝕜] E' := + f.isometryOfInner fun x y => by + rw [← LinearEquiv.coe_coe] at hf + classical rw [← v.linearCombination_repr x, ← v.linearCombination_repr y, + ← LinearEquiv.coe_coe f, Finsupp.apply_linearCombination, + Finsupp.apply_linearCombination, hv.inner_finsupp_eq_sum_left, hf.inner_finsupp_eq_sum_left] + +@[simp] +theorem LinearEquiv.coe_isometryOfOrthonormal (f : E ≃ₗ[𝕜] E') {v : Basis ι 𝕜 E} + (hv : Orthonormal 𝕜 v) (hf : Orthonormal 𝕜 (f ∘ v)) : ⇑(f.isometryOfOrthonormal hv hf) = f := + rfl + +@[simp] +theorem LinearEquiv.isometryOfOrthonormal_toLinearEquiv (f : E ≃ₗ[𝕜] E') {v : Basis ι 𝕜 E} + (hv : Orthonormal 𝕜 v) (hf : Orthonormal 𝕜 (f ∘ v)) : + (f.isometryOfOrthonormal hv hf).toLinearEquiv = f := + rfl + +/-- A linear isometric equivalence that sends an orthonormal basis to a given orthonormal basis. -/ +def Orthonormal.equiv {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) {v' : Basis ι' 𝕜 E'} + (hv' : Orthonormal 𝕜 v') (e : ι ≃ ι') : E ≃ₗᵢ[𝕜] E' := + (v.equiv v' e).isometryOfOrthonormal hv + (by + have h : v.equiv v' e ∘ v = v' ∘ e := by + ext i + simp + rw [h] + classical exact hv'.comp _ e.injective) + +@[simp] +theorem Orthonormal.equiv_toLinearEquiv {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) + {v' : Basis ι' 𝕜 E'} (hv' : Orthonormal 𝕜 v') (e : ι ≃ ι') : + (hv.equiv hv' e).toLinearEquiv = v.equiv v' e := + rfl + +@[simp] +theorem Orthonormal.equiv_apply {ι' : Type*} {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) + {v' : Basis ι' 𝕜 E'} (hv' : Orthonormal 𝕜 v') (e : ι ≃ ι') (i : ι) : + hv.equiv hv' e (v i) = v' (e i) := + Basis.equiv_apply _ _ _ _ + +@[simp] +theorem Orthonormal.equiv_trans {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) {v' : Basis ι' 𝕜 E'} + (hv' : Orthonormal 𝕜 v') (e : ι ≃ ι') {v'' : Basis ι'' 𝕜 E''} (hv'' : Orthonormal 𝕜 v'') + (e' : ι' ≃ ι'') : (hv.equiv hv' e).trans (hv'.equiv hv'' e') = hv.equiv hv'' (e.trans e') := + v.ext_linearIsometryEquiv fun i => by + simp only [LinearIsometryEquiv.trans_apply, Orthonormal.equiv_apply, e.coe_trans, + Function.comp_apply] + +theorem Orthonormal.map_equiv {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) {v' : Basis ι' 𝕜 E'} + (hv' : Orthonormal 𝕜 v') (e : ι ≃ ι') : + v.map (hv.equiv hv' e).toLinearEquiv = v'.reindex e.symm := + v.map_equiv _ _ + +end + +section + +variable {ι : Type*} {ι' : Type*} {E' : Type*} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] + +@[simp] +theorem Orthonormal.equiv_refl {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) : + hv.equiv hv (Equiv.refl ι) = LinearIsometryEquiv.refl 𝕜 E := + v.ext_linearIsometryEquiv fun i => by + simp only [Orthonormal.equiv_apply, Equiv.coe_refl, id, LinearIsometryEquiv.coe_refl] + +@[simp] +theorem Orthonormal.equiv_symm {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) {v' : Basis ι' 𝕜 E'} + (hv' : Orthonormal 𝕜 v') (e : ι ≃ ι') : (hv.equiv hv' e).symm = hv'.equiv hv e.symm := + v'.ext_linearIsometryEquiv fun i => + (hv.equiv hv' e).injective <| by + simp only [LinearIsometryEquiv.apply_symm_apply, Orthonormal.equiv_apply, e.apply_symm_apply] + +end + +end Norm_Seminormed + +section BesselsInequality + +variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] + +variable {ι : Type*} (x : E) {v : ι → E} + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + +/-- Bessel's inequality for finite sums. -/ +theorem Orthonormal.sum_inner_products_le {s : Finset ι} (hv : Orthonormal 𝕜 v) : + ∑ i ∈ s, ‖⟪v i, x⟫‖ ^ 2 ≤ ‖x‖ ^ 2 := by + have h₂ : + (∑ i ∈ s, ∑ j ∈ s, ⟪v i, x⟫ * ⟪x, v j⟫ * ⟪v j, v i⟫) = (∑ k ∈ s, ⟪v k, x⟫ * ⟪x, v k⟫ : 𝕜) := by + classical exact hv.inner_left_right_finset + have h₃ : ∀ z : 𝕜, re (z * conj z) = ‖z‖ ^ 2 := by + intro z + simp only [mul_conj, normSq_eq_def'] + norm_cast + suffices hbf : ‖x - ∑ i ∈ s, ⟪v i, x⟫ • v i‖ ^ 2 = ‖x‖ ^ 2 - ∑ i ∈ s, ‖⟪v i, x⟫‖ ^ 2 by + rw [← sub_nonneg, ← hbf] + simp only [norm_nonneg, pow_nonneg] + rw [@norm_sub_sq 𝕜, sub_add] + simp only [@InnerProductSpace.norm_sq_eq_inner 𝕜 E, inner_sum, sum_inner] + simp only [inner_smul_right, two_mul, inner_smul_left, inner_conj_symm, ← mul_assoc, h₂, + add_sub_cancel_right, sub_right_inj] + simp only [map_sum, ← inner_conj_symm x, ← h₃] + +/-- Bessel's inequality. -/ +theorem Orthonormal.tsum_inner_products_le (hv : Orthonormal 𝕜 v) : + ∑' i, ‖⟪v i, x⟫‖ ^ 2 ≤ ‖x‖ ^ 2 := by + refine tsum_le_of_sum_le' ?_ fun s => hv.sum_inner_products_le x + simp only [norm_nonneg, pow_nonneg] + +/-- The sum defined in Bessel's inequality is summable. -/ +theorem Orthonormal.inner_products_summable (hv : Orthonormal 𝕜 v) : + Summable fun i => ‖⟪v i, x⟫‖ ^ 2 := by + use ⨆ s : Finset ι, ∑ i ∈ s, ‖⟪v i, x⟫‖ ^ 2 + apply hasSum_of_isLUB_of_nonneg + · intro b + simp only [norm_nonneg, pow_nonneg] + · refine isLUB_ciSup ?_ + use ‖x‖ ^ 2 + rintro y ⟨s, rfl⟩ + exact hv.sum_inner_products_le x + +end BesselsInequality diff --git a/Mathlib/Analysis/InnerProductSpace/Subspace.lean b/Mathlib/Analysis/InnerProductSpace/Subspace.lean new file mode 100644 index 0000000000000..a784e77ebd220 --- /dev/null +++ b/Mathlib/Analysis/InnerProductSpace/Subspace.lean @@ -0,0 +1,275 @@ +/- +Copyright (c) 2019 Zhouhang Zhou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Zhouhang Zhou, Sébastien Gouëzel, Frédéric Dupuis +-/ +import Mathlib.Analysis.InnerProductSpace.Orthonormal + +/-! +# Subspaces of inner product spaces + +This file defines the inner-product structure on a subspace of an inner-product space, and proves +some theorems about orthogonal families of subspaces. +-/ + +noncomputable section + +open RCLike Real Filter Topology ComplexConjugate Finsupp + +open LinearMap (BilinForm) + +variable {𝕜 E F : Type*} [RCLike 𝕜] + +section Submodule + +variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + +/-! ### Inner product space structure on subspaces -/ + +/-- Induced inner product on a submodule. -/ +instance Submodule.innerProductSpace (W : Submodule 𝕜 E) : InnerProductSpace 𝕜 W := + { Submodule.normedSpace W with + inner := fun x y => ⟪(x : E), (y : E)⟫ + conj_symm := fun _ _ => inner_conj_symm _ _ + norm_sq_eq_inner := fun x => norm_sq_eq_inner (x : E) + add_left := fun _ _ _ => inner_add_left _ _ _ + smul_left := fun _ _ _ => inner_smul_left _ _ _ } + +/-- The inner product on submodules is the same as on the ambient space. -/ +@[simp] +theorem Submodule.coe_inner (W : Submodule 𝕜 E) (x y : W) : ⟪x, y⟫ = ⟪(x : E), ↑y⟫ := + rfl + +theorem Orthonormal.codRestrict {ι : Type*} {v : ι → E} (hv : Orthonormal 𝕜 v) (s : Submodule 𝕜 E) + (hvs : ∀ i, v i ∈ s) : @Orthonormal 𝕜 s _ _ _ ι (Set.codRestrict v s hvs) := + s.subtypeₗᵢ.orthonormal_comp_iff.mp hv + +theorem orthonormal_span {ι : Type*} {v : ι → E} (hv : Orthonormal 𝕜 v) : + @Orthonormal 𝕜 (Submodule.span 𝕜 (Set.range v)) _ _ _ ι fun i : ι => + ⟨v i, Submodule.subset_span (Set.mem_range_self i)⟩ := + hv.codRestrict (Submodule.span 𝕜 (Set.range v)) fun i => + Submodule.subset_span (Set.mem_range_self i) + +end Submodule + +/-! ### Families of mutually-orthogonal subspaces of an inner product space -/ + +section OrthogonalFamily_Seminormed + +variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + +variable {ι : Type*} (𝕜) + +open DirectSum + +/-- An indexed family of mutually-orthogonal subspaces of an inner product space `E`. + +The simple way to express this concept would be as a condition on `V : ι → Submodule 𝕜 E`. We +instead implement it as a condition on a family of inner product spaces each equipped with an +isometric embedding into `E`, thus making it a property of morphisms rather than subobjects. +The connection to the subobject spelling is shown in `orthogonalFamily_iff_pairwise`. + +This definition is less lightweight, but allows for better definitional properties when the inner +product space structure on each of the submodules is important -- for example, when considering +their Hilbert sum (`PiLp V 2`). For example, given an orthonormal set of vectors `v : ι → E`, +we have an associated orthogonal family of one-dimensional subspaces of `E`, which it is convenient +to be able to discuss using `ι → 𝕜` rather than `Π i : ι, span 𝕜 (v i)`. -/ +def OrthogonalFamily (G : ι → Type*) [∀ i, SeminormedAddCommGroup (G i)] + [∀ i, InnerProductSpace 𝕜 (G i)] (V : ∀ i, G i →ₗᵢ[𝕜] E) : Prop := + Pairwise fun i j => ∀ v : G i, ∀ w : G j, ⟪V i v, V j w⟫ = 0 + +variable {𝕜} +variable {G : ι → Type*} [∀ i, NormedAddCommGroup (G i)] [∀ i, InnerProductSpace 𝕜 (G i)] + {V : ∀ i, G i →ₗᵢ[𝕜] E} + +theorem Orthonormal.orthogonalFamily {v : ι → E} (hv : Orthonormal 𝕜 v) : + OrthogonalFamily 𝕜 (fun _i : ι => 𝕜) fun i => LinearIsometry.toSpanSingleton 𝕜 E (hv.1 i) := + fun i j hij a b => by simp [inner_smul_left, inner_smul_right, hv.2 hij] + +section +variable (hV : OrthogonalFamily 𝕜 G V) +include hV + +theorem OrthogonalFamily.eq_ite [DecidableEq ι] {i j : ι} (v : G i) (w : G j) : + ⟪V i v, V j w⟫ = ite (i = j) ⟪V i v, V j w⟫ 0 := by + split_ifs with h + · rfl + · exact hV h v w + +theorem OrthogonalFamily.inner_right_dfinsupp + [∀ (i) (x : G i), Decidable (x ≠ 0)] [DecidableEq ι] (l : ⨁ i, G i) (i : ι) (v : G i) : + ⟪V i v, l.sum fun j => V j⟫ = ⟪v, l i⟫ := + calc + ⟪V i v, l.sum fun j => V j⟫ = l.sum fun j => fun w => ⟪V i v, V j w⟫ := + DFinsupp.inner_sum (fun j => V j) l (V i v) + _ = l.sum fun j => fun w => ite (i = j) ⟪V i v, V j w⟫ 0 := + (congr_arg l.sum <| funext fun _ => funext <| hV.eq_ite v) + _ = ⟪v, l i⟫ := by + simp only [DFinsupp.sum, Submodule.coe_inner, Finset.sum_ite_eq, ite_eq_left_iff, + DFinsupp.mem_support_toFun] + split_ifs with h + · simp only [LinearIsometry.inner_map_map] + · simp only [of_not_not h, inner_zero_right] + +theorem OrthogonalFamily.inner_right_fintype [Fintype ι] (l : ∀ i, G i) (i : ι) (v : G i) : + ⟪V i v, ∑ j : ι, V j (l j)⟫ = ⟪v, l i⟫ := by + classical + calc + ⟪V i v, ∑ j : ι, V j (l j)⟫ = ∑ j : ι, ⟪V i v, V j (l j)⟫ := by rw [inner_sum] + _ = ∑ j, ite (i = j) ⟪V i v, V j (l j)⟫ 0 := + (congr_arg (Finset.sum Finset.univ) <| funext fun j => hV.eq_ite v (l j)) + _ = ⟪v, l i⟫ := by + simp only [Finset.sum_ite_eq, Finset.mem_univ, (V i).inner_map_map, if_true] + +nonrec theorem OrthogonalFamily.inner_sum (l₁ l₂ : ∀ i, G i) (s : Finset ι) : + ⟪∑ i ∈ s, V i (l₁ i), ∑ j ∈ s, V j (l₂ j)⟫ = ∑ i ∈ s, ⟪l₁ i, l₂ i⟫ := by + classical + calc + ⟪∑ i ∈ s, V i (l₁ i), ∑ j ∈ s, V j (l₂ j)⟫ = ∑ j ∈ s, ∑ i ∈ s, ⟪V i (l₁ i), V j (l₂ j)⟫ := by + simp only [sum_inner, inner_sum] + _ = ∑ j ∈ s, ∑ i ∈ s, ite (i = j) ⟪V i (l₁ i), V j (l₂ j)⟫ 0 := by + congr with i + congr with j + apply hV.eq_ite + _ = ∑ i ∈ s, ⟪l₁ i, l₂ i⟫ := by + simp only [Finset.sum_ite_of_true, Finset.sum_ite_eq', LinearIsometry.inner_map_map, + imp_self, imp_true_iff] + +theorem OrthogonalFamily.norm_sum (l : ∀ i, G i) (s : Finset ι) : + ‖∑ i ∈ s, V i (l i)‖ ^ 2 = ∑ i ∈ s, ‖l i‖ ^ 2 := by + have : ((‖∑ i ∈ s, V i (l i)‖ : ℝ) : 𝕜) ^ 2 = ∑ i ∈ s, ((‖l i‖ : ℝ) : 𝕜) ^ 2 := by + simp only [← inner_self_eq_norm_sq_to_K, hV.inner_sum] + exact mod_cast this + +/-- The composition of an orthogonal family of subspaces with an injective function is also an +orthogonal family. -/ +theorem OrthogonalFamily.comp {γ : Type*} {f : γ → ι} (hf : Function.Injective f) : + OrthogonalFamily 𝕜 (fun g => G (f g)) fun g => V (f g) := + fun _i _j hij v w => hV (hf.ne hij) v w + +theorem OrthogonalFamily.orthonormal_sigma_orthonormal {α : ι → Type*} {v_family : ∀ i, α i → G i} + (hv_family : ∀ i, Orthonormal 𝕜 (v_family i)) : + Orthonormal 𝕜 fun a : Σi, α i => V a.1 (v_family a.1 a.2) := by + constructor + · rintro ⟨i, v⟩ + simpa only [LinearIsometry.norm_map] using (hv_family i).left v + rintro ⟨i, v⟩ ⟨j, w⟩ hvw + by_cases hij : i = j + · subst hij + have : v ≠ w := fun h => by + subst h + exact hvw rfl + simpa only [LinearIsometry.inner_map_map] using (hv_family i).2 this + · exact hV hij (v_family i v) (v_family j w) + +theorem OrthogonalFamily.norm_sq_diff_sum [DecidableEq ι] (f : ∀ i, G i) (s₁ s₂ : Finset ι) : + ‖(∑ i ∈ s₁, V i (f i)) - ∑ i ∈ s₂, V i (f i)‖ ^ 2 = + (∑ i ∈ s₁ \ s₂, ‖f i‖ ^ 2) + ∑ i ∈ s₂ \ s₁, ‖f i‖ ^ 2 := by + rw [← Finset.sum_sdiff_sub_sum_sdiff, sub_eq_add_neg, ← Finset.sum_neg_distrib] + let F : ∀ i, G i := fun i => if i ∈ s₁ then f i else -f i + have hF₁ : ∀ i ∈ s₁ \ s₂, F i = f i := fun i hi => if_pos (Finset.sdiff_subset hi) + have hF₂ : ∀ i ∈ s₂ \ s₁, F i = -f i := fun i hi => if_neg (Finset.mem_sdiff.mp hi).2 + have hF : ∀ i, ‖F i‖ = ‖f i‖ := by + intro i + dsimp only [F] + split_ifs <;> simp only [eq_self_iff_true, norm_neg] + have : + ‖(∑ i ∈ s₁ \ s₂, V i (F i)) + ∑ i ∈ s₂ \ s₁, V i (F i)‖ ^ 2 = + (∑ i ∈ s₁ \ s₂, ‖F i‖ ^ 2) + ∑ i ∈ s₂ \ s₁, ‖F i‖ ^ 2 := by + have hs : Disjoint (s₁ \ s₂) (s₂ \ s₁) := disjoint_sdiff_sdiff + simpa only [Finset.sum_union hs] using hV.norm_sum F (s₁ \ s₂ ∪ s₂ \ s₁) + convert this using 4 + · refine Finset.sum_congr rfl fun i hi => ?_ + simp only [hF₁ i hi] + · refine Finset.sum_congr rfl fun i hi => ?_ + simp only [hF₂ i hi, LinearIsometry.map_neg] + · simp only [hF] + · simp only [hF] + +/-- A family `f` of mutually-orthogonal elements of `E` is summable, if and only if +`(fun i ↦ ‖f i‖ ^ 2)` is summable. -/ +theorem OrthogonalFamily.summable_iff_norm_sq_summable [CompleteSpace E] (f : ∀ i, G i) : + (Summable fun i => V i (f i)) ↔ Summable fun i => ‖f i‖ ^ 2 := by + classical + simp only [summable_iff_cauchySeq_finset, NormedAddCommGroup.cauchySeq_iff, Real.norm_eq_abs] + constructor + · intro hf ε hε + obtain ⟨a, H⟩ := hf _ (sqrt_pos.mpr hε) + use a + intro s₁ hs₁ s₂ hs₂ + rw [← Finset.sum_sdiff_sub_sum_sdiff] + refine (abs_sub _ _).trans_lt ?_ + have : ∀ i, 0 ≤ ‖f i‖ ^ 2 := fun i : ι => sq_nonneg _ + simp only [Finset.abs_sum_of_nonneg' this] + have : ((∑ i ∈ s₁ \ s₂, ‖f i‖ ^ 2) + ∑ i ∈ s₂ \ s₁, ‖f i‖ ^ 2) < √ε ^ 2 := by + rw [← hV.norm_sq_diff_sum, sq_lt_sq, abs_of_nonneg (sqrt_nonneg _), + abs_of_nonneg (norm_nonneg _)] + exact H s₁ hs₁ s₂ hs₂ + have hη := sq_sqrt (le_of_lt hε) + linarith + · intro hf ε hε + have hε' : 0 < ε ^ 2 / 2 := half_pos (sq_pos_of_pos hε) + obtain ⟨a, H⟩ := hf _ hε' + use a + intro s₁ hs₁ s₂ hs₂ + refine (abs_lt_of_sq_lt_sq' ?_ (le_of_lt hε)).2 + have has : a ≤ s₁ ⊓ s₂ := le_inf hs₁ hs₂ + rw [hV.norm_sq_diff_sum] + have Hs₁ : ∑ x ∈ s₁ \ s₂, ‖f x‖ ^ 2 < ε ^ 2 / 2 := by + convert H _ hs₁ _ has + have : s₁ ⊓ s₂ ⊆ s₁ := Finset.inter_subset_left + rw [← Finset.sum_sdiff this, add_tsub_cancel_right, Finset.abs_sum_of_nonneg'] + · simp + · exact fun i => sq_nonneg _ + have Hs₂ : ∑ x ∈ s₂ \ s₁, ‖f x‖ ^ 2 < ε ^ 2 / 2 := by + convert H _ hs₂ _ has + have : s₁ ⊓ s₂ ⊆ s₂ := Finset.inter_subset_right + rw [← Finset.sum_sdiff this, add_tsub_cancel_right, Finset.abs_sum_of_nonneg'] + · simp + · exact fun i => sq_nonneg _ + linarith + +end + +end OrthogonalFamily_Seminormed + +section OrthogonalFamily + +variable [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + +variable {ι : Type*} {G : ι → Type*} + +/-- An orthogonal family forms an independent family of subspaces; that is, any collection of +elements each from a different subspace in the family is linearly independent. In particular, the +pairwise intersections of elements of the family are 0. -/ +theorem OrthogonalFamily.independent {V : ι → Submodule 𝕜 E} + (hV : OrthogonalFamily 𝕜 (fun i => V i) fun i => (V i).subtypeₗᵢ) : + iSupIndep V := by + classical + apply iSupIndep_of_dfinsupp_lsum_injective + refine LinearMap.ker_eq_bot.mp ?_ + rw [Submodule.eq_bot_iff] + intro v hv + rw [LinearMap.mem_ker] at hv + ext i + suffices ⟪(v i : E), v i⟫ = 0 by simpa only [inner_self_eq_zero] using this + calc + ⟪(v i : E), v i⟫ = ⟪(v i : E), DFinsupp.lsum ℕ (fun i => (V i).subtype) v⟫ := by + simpa only [DFinsupp.sumAddHom_apply, DFinsupp.lsum_apply_apply] using + (hV.inner_right_dfinsupp v i (v i)).symm + _ = 0 := by simp only [hv, inner_zero_right] + +theorem DirectSum.IsInternal.collectedBasis_orthonormal [DecidableEq ι] {V : ι → Submodule 𝕜 E} + (hV : OrthogonalFamily 𝕜 (fun i => V i) fun i => (V i).subtypeₗᵢ) + (hV_sum : DirectSum.IsInternal fun i => V i) {α : ι → Type*} + {v_family : ∀ i, Basis (α i) 𝕜 (V i)} (hv_family : ∀ i, Orthonormal 𝕜 (v_family i)) : + Orthonormal 𝕜 (hV_sum.collectedBasis v_family) := by + simpa only [hV_sum.collectedBasis_coe] using hV.orthonormal_sigma_orthonormal hv_family + +end OrthogonalFamily diff --git a/Mathlib/Analysis/InnerProductSpace/Symmetric.lean b/Mathlib/Analysis/InnerProductSpace/Symmetric.lean index fe6587c770dc5..0e01e99c76656 100644 --- a/Mathlib/Analysis/InnerProductSpace/Symmetric.lean +++ b/Mathlib/Analysis/InnerProductSpace/Symmetric.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Anatole Dedecker. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Moritz Doll, Frédéric Dupuis, Heather Macbeth -/ -import Mathlib.Analysis.InnerProductSpace.Basic +import Mathlib.Analysis.InnerProductSpace.Subspace import Mathlib.Analysis.Normed.Operator.Banach import Mathlib.LinearAlgebra.SesquilinearForm diff --git a/Mathlib/Geometry/Euclidean/Angle/Unoriented/Basic.lean b/Mathlib/Geometry/Euclidean/Angle/Unoriented/Basic.lean index 4b5fc638c7a20..804654e42dff5 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Unoriented/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Unoriented/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Joseph Myers. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers, Manuel Candales -/ -import Mathlib.Analysis.InnerProductSpace.Basic +import Mathlib.Analysis.InnerProductSpace.Subspace import Mathlib.Analysis.SpecialFunctions.Trigonometric.Inverse /-! diff --git a/Mathlib/MeasureTheory/Function/L2Space.lean b/Mathlib/MeasureTheory/Function/L2Space.lean index e89d919547be1..de18612ca8356 100644 --- a/Mathlib/MeasureTheory/Function/L2Space.lean +++ b/Mathlib/MeasureTheory/Function/L2Space.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Analysis.RCLike.Lemmas +import Mathlib.Analysis.InnerProductSpace.LinearMap import Mathlib.MeasureTheory.Function.StronglyMeasurable.Inner import Mathlib.MeasureTheory.Integral.SetIntegral diff --git a/scripts/noshake.json b/scripts/noshake.json index 3ae49049109e5..6339d0b82139c 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -425,8 +425,7 @@ "Mathlib.Analysis.Normed.Operator.LinearIsometry": ["Mathlib.Algebra.Star.Basic"], "Mathlib.Analysis.InnerProductSpace.PiL2": ["Mathlib.Util.Superscript"], - "Mathlib.Analysis.InnerProductSpace.Basic": - ["Mathlib.Algebra.Module.LinearMap.Basic"], + "Mathlib.Analysis.InnerProductSpace.Defs": ["Mathlib.Data.Complex.Basic"], "Mathlib.Analysis.Distribution.SchwartzSpace": ["Mathlib.Tactic.MoveAdd"], "Mathlib.Analysis.Convex.Star": ["Mathlib.Algebra.Order.Module.Synonym"], "Mathlib.Analysis.Convex.Basic": From fb92079b5b94ef215000eef5e849b6be4d3c6e5d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 20 Dec 2024 05:04:01 +0000 Subject: [PATCH 814/829] feat(Finset): `(s \ t).image f = s.image f \ t.image f` when `f` is injective on `s` (#20057) From LeanCamCombi --- Mathlib/Data/Finset/Image.lean | 4 ++++ Mathlib/Data/Set/Function.lean | 2 ++ 2 files changed, 6 insertions(+) diff --git a/Mathlib/Data/Finset/Image.lean b/Mathlib/Data/Finset/Image.lean index f5732b4cbc6e6..0efb04b75e74c 100644 --- a/Mathlib/Data/Finset/Image.lean +++ b/Mathlib/Data/Finset/Image.lean @@ -483,6 +483,10 @@ theorem image_sdiff [DecidableEq α] {f : α → β} (s t : Finset α) (hf : Inj (s \ t).image f = s.image f \ t.image f := mod_cast Set.image_diff hf s t +lemma image_sdiff_of_injOn [DecidableEq α] {t : Finset α} (hf : Set.InjOn f s) (hts : t ⊆ s) : + (s \ t).image f = s.image f \ t.image f := + mod_cast Set.image_diff_of_injOn hf <| coe_subset.2 hts + open scoped symmDiff in theorem image_symmDiff [DecidableEq α] {f : α → β} (s t : Finset α) (hf : Injective f) : (s ∆ t).image f = s.image f ∆ t.image f := diff --git a/Mathlib/Data/Set/Function.lean b/Mathlib/Data/Set/Function.lean index 77dbaaa172001..f2c6f67261d56 100644 --- a/Mathlib/Data/Set/Function.lean +++ b/Mathlib/Data/Set/Function.lean @@ -610,6 +610,8 @@ lemma InjOn.image_diff_subset {f : α → β} {t : Set α} (h : InjOn f s) (hst f '' (s \ t) = f '' s \ f '' t := by rw [h.image_diff, inter_eq_self_of_subset_right hst] +alias image_diff_of_injOn := InjOn.image_diff_subset + theorem InjOn.imageFactorization_injective (h : InjOn f s) : Injective (s.imageFactorization f) := fun ⟨x, hx⟩ ⟨y, hy⟩ h' ↦ by simpa [imageFactorization, h.eq_iff hx hy] using h' From 990a16dffbf6acf2e536247009a6826284ef0379 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 20 Dec 2024 05:04:02 +0000 Subject: [PATCH 815/829] =?UTF-8?q?feat:=20`Coprime=20(a=20+=20b)=20c=20?= =?UTF-8?q?=E2=86=94=20Coprime=20a=20c`=20if=20`c=20=E2=88=A3=20b`=20(#200?= =?UTF-8?q?60)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From LeanCamCombi --- Mathlib/Data/Nat/GCD/Basic.lean | 23 +++++++++++++++++++++++ 1 file changed, 23 insertions(+) diff --git a/Mathlib/Data/Nat/GCD/Basic.lean b/Mathlib/Data/Nat/GCD/Basic.lean index 78cfdec77c654..30acba25218bd 100644 --- a/Mathlib/Data/Nat/GCD/Basic.lean +++ b/Mathlib/Data/Nat/GCD/Basic.lean @@ -25,6 +25,7 @@ Most of this file could be moved to batteries as well. assert_not_exists OrderedCommMonoid namespace Nat +variable {a a₁ a₂ b b₁ b₂ c : ℕ} /-! ### `gcd` -/ @@ -193,6 +194,28 @@ theorem coprime_mul_right_add_left (m n k : ℕ) : Coprime (k * n + m) n ↔ Cop theorem coprime_mul_left_add_left (m n k : ℕ) : Coprime (n * k + m) n ↔ Coprime m n := by rw [Coprime, Coprime, gcd_mul_left_add_left] +lemma add_coprime_iff_left (h : c ∣ b) : Coprime (a + b) c ↔ Coprime a c := by + obtain ⟨n, rfl⟩ := h; simp + +lemma add_coprime_iff_right (h : c ∣ a) : Coprime (a + b) c ↔ Coprime b c := by + obtain ⟨n, rfl⟩ := h; simp + +lemma coprime_add_iff_left (h : a ∣ c) : Coprime a (b + c) ↔ Coprime a b := by + obtain ⟨n, rfl⟩ := h; simp + +lemma coprime_add_iff_right (h : a ∣ b) : Coprime a (b + c) ↔ Coprime a c := by + obtain ⟨n, rfl⟩ := h; simp + +-- TODO: Replace `Nat.Coprime.coprime_dvd_left` +lemma Coprime.of_dvd_left (ha : a₁ ∣ a₂) (h : Coprime a₂ b) : Coprime a₁ b := h.coprime_dvd_left ha + +-- TODO: Replace `Nat.Coprime.coprime_dvd_right` +lemma Coprime.of_dvd_right (hb : b₁ ∣ b₂) (h : Coprime a b₂) : Coprime a b₁ := + h.coprime_dvd_right hb + +lemma Coprime.of_dvd (ha : a₁ ∣ a₂) (hb : b₁ ∣ b₂) (h : Coprime a₂ b₂) : Coprime a₁ b₁ := + (h.of_dvd_left ha).of_dvd_right hb + @[simp] theorem coprime_sub_self_left {m n : ℕ} (h : m ≤ n) : Coprime (n - m) m ↔ Coprime n m := by rw [Coprime, Coprime, gcd_sub_self_left h] From a81f5e199618895d0c377524ed5915038c88523a Mon Sep 17 00:00:00 2001 From: Rida Hamadani Date: Fri, 20 Dec 2024 05:04:03 +0000 Subject: [PATCH 816/829] feat(Data/Matroid): `disjointSum` is commutative (#20074) from the [Seymour](https://github.com/Ivan-Sergeyev/seymour) project. --- Mathlib/Data/Matroid/Sum.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/Mathlib/Data/Matroid/Sum.lean b/Mathlib/Data/Matroid/Sum.lean index 77f18e0275913..f3e98ee11cd2c 100644 --- a/Mathlib/Data/Matroid/Sum.lean +++ b/Mathlib/Data/Matroid/Sum.lean @@ -276,6 +276,12 @@ def disjointSum (M N : Matroid α) (h : Disjoint M.E N.E) : Matroid α := N.Basis (I ∩ N.E) (X ∩ N.E) ∧ I ⊆ X ∧ X ⊆ M.E ∪ N.E := by simp [disjointSum, and_assoc] +lemma disjointSum_comm {h} : M.disjointSum N h = N.disjointSum M h.symm := by + ext + · simp [union_comm] + repeat simpa [union_comm] using ⟨fun ⟨m, n, h⟩ ↦ ⟨n, m, M.E.union_comm N.E ▸ h⟩, + fun ⟨n, m, h⟩ ↦ ⟨m, n, M.E.union_comm N.E ▸ h⟩⟩ + lemma Indep.eq_union_image_of_disjointSum {h I} (hI : (disjointSum M N h).Indep I) : ∃ IM IN, M.Indep IM ∧ N.Indep IN ∧ Disjoint IM IN ∧ I = IM ∪ IN := by rw [disjointSum_indep_iff] at hI From ea22acdd1b6522affd5d2249836afa9d18ca59f8 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 20 Dec 2024 05:45:11 +0000 Subject: [PATCH 817/829] feat: add bundled multiplication maps for matrices (#19229) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This repeats the definitions and theorems about `LinearMap.mulLeft` and `LinearMap.mulRight` for rectangular matrices. These definitions (or rather, the continuous versions of them) are needed to state results about Fréchet derivatives of matrices. --- Mathlib.lean | 1 + Mathlib/Data/Matrix/Basis.lean | 18 +-- Mathlib/Data/Matrix/Bilinear.lean | 192 ++++++++++++++++++++++++++++++ 3 files changed, 204 insertions(+), 7 deletions(-) create mode 100644 Mathlib/Data/Matrix/Bilinear.lean diff --git a/Mathlib.lean b/Mathlib.lean index 5f0ed064cabd5..a43ecaafda619 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2644,6 +2644,7 @@ import Mathlib.Data.MLList.Split import Mathlib.Data.Matrix.Auto import Mathlib.Data.Matrix.Basic import Mathlib.Data.Matrix.Basis +import Mathlib.Data.Matrix.Bilinear import Mathlib.Data.Matrix.Block import Mathlib.Data.Matrix.CharP import Mathlib.Data.Matrix.ColumnRowPartitioned diff --git a/Mathlib/Data/Matrix/Basis.lean b/Mathlib/Data/Matrix/Basis.lean index 7c49644844bcd..69638ec2666c4 100644 --- a/Mathlib/Data/Matrix/Basis.lean +++ b/Mathlib/Data/Matrix/Basis.lean @@ -153,33 +153,37 @@ theorem diag_same : diag (stdBasisMatrix i i c) = Pi.single i c := by end section mul -variable [Fintype n] [NonUnitalNonAssocSemiring α] (i j : n) (c : α) +variable [Fintype m] [NonUnitalNonAssocSemiring α] (c : α) +omit [DecidableEq n] in @[simp] -theorem mul_left_apply_same (b : n) (M : Matrix n n α) : +theorem mul_left_apply_same (i : l) (j : m) (b : n) (M : Matrix m n α) : (stdBasisMatrix i j c * M) i b = c * M j b := by simp [mul_apply, stdBasisMatrix] +omit [DecidableEq l] in @[simp] -theorem mul_right_apply_same (a : n) (M : Matrix n n α) : +theorem mul_right_apply_same (i : m) (j : n) (a : l) (M : Matrix l m α) : (M * stdBasisMatrix i j c) a j = M a i * c := by simp [mul_apply, stdBasisMatrix, mul_comm] +omit [DecidableEq n] in @[simp] -theorem mul_left_apply_of_ne (a b : n) (h : a ≠ i) (M : Matrix n n α) : +theorem mul_left_apply_of_ne (i : l) (j : m) (a : l) (b : n) (h : a ≠ i) (M : Matrix m n α) : (stdBasisMatrix i j c * M) a b = 0 := by simp [mul_apply, h.symm] +omit [DecidableEq l] in @[simp] -theorem mul_right_apply_of_ne (a b : n) (hbj : b ≠ j) (M : Matrix n n α) : +theorem mul_right_apply_of_ne (i : m) (j : n) (a : l) (b : n) (hbj : b ≠ j) (M : Matrix l m α) : (M * stdBasisMatrix i j c) a b = 0 := by simp [mul_apply, hbj.symm] @[simp] -theorem mul_same (k : n) (d : α) : +theorem mul_same (i : l) (j : m) (k : n) (d : α) : stdBasisMatrix i j c * stdBasisMatrix j k d = stdBasisMatrix i k (c * d) := by ext a b simp only [mul_apply, stdBasisMatrix, boole_mul] by_cases h₁ : i = a <;> by_cases h₂ : k = b <;> simp [h₁, h₂] @[simp] -theorem mul_of_ne {k l : n} (h : j ≠ k) (d : α) : +theorem mul_of_ne (i : l) (j k : m) {l : n} (h : j ≠ k) (d : α) : stdBasisMatrix i j c * stdBasisMatrix k l d = 0 := by ext a b simp only [mul_apply, boole_mul, stdBasisMatrix, of_apply] diff --git a/Mathlib/Data/Matrix/Bilinear.lean b/Mathlib/Data/Matrix/Bilinear.lean new file mode 100644 index 0000000000000..9d34a4087f577 --- /dev/null +++ b/Mathlib/Data/Matrix/Bilinear.lean @@ -0,0 +1,192 @@ +/- +Copyright (c) 2024 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import Mathlib.Algebra.Module.LinearMap.End +import Mathlib.Data.Matrix.Mul +import Mathlib.Data.Matrix.Basis +import Mathlib.Algebra.Algebra.Bilinear + +/-! +# Bundled versions of multiplication for matrices + +This file provides versions of `LinearMap.mulLeft` and `LinearMap.mulRight` which work for the +heterogeneous multiplication of matrices. +-/ + +variable {l m n o : Type*} {R A : Type*} + +section NonUnitalNonAssocSemiring +variable (R) [Fintype m] + +section one_side +variable [Semiring R] [NonUnitalNonAssocSemiring A] [Module R A] + +section left +variable (n) [SMulCommClass R A A] + +/-- A version of `LinearMap.mulLeft` for matrix multiplication. -/ +@[simps] +def mulLeftLinearMap (X : Matrix l m A) : + Matrix m n A →ₗ[R] Matrix l n A where + toFun := (X * ·) + map_smul' := Matrix.mul_smul _ + map_add' := Matrix.mul_add _ + +/-- On square matrices, `Matrix.mulLeftLinearMap` and `LinearMap.mulLeft` coincide. -/ +theorem mulLeftLinearMap_eq_mulLeft : + mulLeftLinearMap m R = LinearMap.mulLeft R (A := Matrix m m A) := rfl + +/-- A version of `LinearMap.mulLeft_zero_eq_zero` for matrix multiplication. -/ +@[simp] +theorem mulLeftLinearMap_zero_eq_zero : + mulLeftLinearMap n R (0 : Matrix l m A) = 0 := LinearMap.ext fun _ => Matrix.zero_mul _ + +end left + +section right +variable (l) [IsScalarTower R A A] + +/-- A version of `LinearMap.mulRight` for matrix multiplication. -/ +@[simps] +def mulRightLinearMap (Y : Matrix m n A) : + Matrix l m A →ₗ[R] Matrix l n A where + toFun := (· * Y) + map_smul' _ _ := Matrix.smul_mul _ _ _ + map_add' _ _ := Matrix.add_mul _ _ _ + +/-- On square matrices, `Matrix.mulRightLinearMap` and `LinearMap.mulRight` coincide. -/ +theorem mulRightLinearMap_eq_mulRight : + mulRightLinearMap m R = LinearMap.mulRight R (A := Matrix m m A) := rfl + +/-- A version of `LinearMap.mulLeft_zero_eq_zero` for matrix multiplication. -/ +@[simp] +theorem mulRightLinearMap_zero_eq_zero : + mulRightLinearMap l R (0 : Matrix m n A) = 0 := LinearMap.ext fun _ => Matrix.mul_zero _ + +end right + +end one_side + +variable [CommSemiring R] [NonUnitalNonAssocSemiring A] [Module R A] +variable [SMulCommClass R A A] [IsScalarTower R A A] + +/-- A version of `LinearMap.mul` for matrix multiplication. -/ +@[simps!] +def mulLinearMap : Matrix l m A →ₗ[R] Matrix m n A →ₗ[R] Matrix l n A where + toFun := mulLeftLinearMap n R + map_add' _ _ := LinearMap.ext fun _ => Matrix.add_mul _ _ _ + map_smul' _ _ := LinearMap.ext fun _ => Matrix.smul_mul _ _ _ + +/-- On square matrices, `Matrix.mulLinearMap` and `LinearMap.mul` coincide. -/ +theorem mulLinearMap_eq_mul : + mulLinearMap R = LinearMap.mul R (A := Matrix m m A) := rfl + +end NonUnitalNonAssocSemiring + +section NonUnital + +section one_side +variable [Fintype m] [Fintype n] [Semiring R] [NonUnitalSemiring A] [Module R A] + +/-- A version of `LinearMap.mulLeft_mul` for matrix multiplication. -/ +@[simp] +theorem mulLeftLinearMap_mul [SMulCommClass R A A] (a : Matrix l m A) (b : Matrix m n A) : + mulLeftLinearMap o R (a * b) = (mulLeftLinearMap o R a).comp (mulLeftLinearMap o R b) := by + ext + simp only [mulLeftLinearMap_apply, LinearMap.comp_apply, Matrix.mul_assoc] + +/-- A version of `LinearMap.mulRight_mul` for matrix multiplication. -/ +@[simp] +theorem mulRightLinearMap_mul [IsScalarTower R A A] (a : Matrix m n A) (b : Matrix n o A) : + mulRightLinearMap l R (a * b) = (mulRightLinearMap l R b).comp (mulRightLinearMap l R a) := by + ext + simp only [mulRightLinearMap_apply, LinearMap.comp_apply, Matrix.mul_assoc] + +end one_side + +variable [Fintype m] [Fintype n] [CommSemiring R] [NonUnitalSemiring A] [Module R A] +variable [SMulCommClass R A A] [IsScalarTower R A A] + +/-- A version of `LinearMap.commute_mulLeft_right` for matrix multiplication. -/ +theorem commute_mulLeftLinearMap_mulRightLinearMap (a : Matrix l m A) (b : Matrix n o A) : + mulLeftLinearMap o R a ∘ₗ mulRightLinearMap m R b = + mulRightLinearMap l R b ∘ₗ mulLeftLinearMap n R a := by + ext c : 1 + exact (Matrix.mul_assoc a c b).symm + +end NonUnital + +section Semiring + +section one_side +variable [Fintype m] [DecidableEq m] [Semiring R] [Semiring A] + +section left +variable [Module R A] [SMulCommClass R A A] + +/-- A version of `LinearMap.mulLeft_one` for matrix multiplication. -/ +@[simp] +theorem mulLeftLinearMap_one : mulLeftLinearMap n R (1 : Matrix m m A) = LinearMap.id := + LinearMap.ext fun _ => Matrix.one_mul _ + +/-- A version of `LinearMap.mulLeft_eq_zero_iff` for matrix multiplication. -/ +@[simp] +theorem mulLeftLinearMap_eq_zero_iff [Nonempty n] (a : Matrix l m A) : + mulLeftLinearMap n R a = 0 ↔ a = 0 := by + constructor <;> intro h + · inhabit n + ext i j + classical + replace h := DFunLike.congr_fun h (Matrix.stdBasisMatrix j (default : n) 1) + simpa using Matrix.ext_iff.2 h i default + · rw [h] + exact mulLeftLinearMap_zero_eq_zero _ _ + +/-- A version of `LinearMap.pow_mulLeft` for matrix multiplication. -/ +@[simp] +theorem pow_mulLeftLinearMap (a : Matrix m m A) (k : ℕ) : + mulLeftLinearMap n R a ^ k = mulLeftLinearMap n R (a ^ k) := + match k with + | 0 => by rw [pow_zero, pow_zero, mulLeftLinearMap_one, LinearMap.one_eq_id] + | (n + 1) => by + rw [pow_succ, pow_succ, mulLeftLinearMap_mul, LinearMap.mul_eq_comp, pow_mulLeftLinearMap] + +end left + +section right +variable [Module R A] [IsScalarTower R A A] + +/-- A version of `LinearMap.mulRight_one` for matrix multiplication. -/ +@[simp] +theorem mulRightLinearMap_one : mulRightLinearMap l R (1 : Matrix m m A) = LinearMap.id := + LinearMap.ext fun _ => Matrix.mul_one _ + +/-- A version of `LinearMap.mulRight_eq_zero_iff` for matrix multiplication. -/ +@[simp] +theorem mulRightLinearMap_eq_zero_iff (a : Matrix m n A) [Nonempty l] : + mulRightLinearMap l R a = 0 ↔ a = 0 := by + constructor <;> intro h + · inhabit l + ext i j + classical + replace h := DFunLike.congr_fun h (Matrix.stdBasisMatrix (default : l) i 1) + simpa using Matrix.ext_iff.2 h default j + · rw [h] + exact mulRightLinearMap_zero_eq_zero _ _ + +/-- A version of `LinearMap.pow_mulRight` for matrix multiplication. -/ +@[simp] +theorem pow_mulRightLinearMap (a : Matrix m m A) (k : ℕ) : + mulRightLinearMap l R a ^ k = mulRightLinearMap l R (a ^ k) := + match k with + | 0 => by rw [pow_zero, pow_zero, mulRightLinearMap_one, LinearMap.one_eq_id] + | (n + 1) => by + rw [pow_succ, pow_succ', mulRightLinearMap_mul, LinearMap.mul_eq_comp, pow_mulRightLinearMap] + +end right + +end one_side + +end Semiring From f4ff33e411c170da90251471bfba29a8486cc831 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Fri, 20 Dec 2024 05:45:13 +0000 Subject: [PATCH 818/829] feat: the second derivative of an analytic map is symmetric (#19770) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Unify the facts that the second derivative of a `C^2` function over R or C, or an analytic function over an arbitrary field, is symmetric. This is expressed as an assumption that `f` is `C^n` for some `n ≥ minSmoothness 𝕜 2`, where `minSmoothness 𝕜 i` is by definition `i` if the field is R or C, and `ω` otherwise. --- .../Analysis/Calculus/FDeriv/Symmetric.lean | 120 ++++++++++++++---- Mathlib/Analysis/Calculus/VectorField.lean | 47 +++++-- 2 files changed, 129 insertions(+), 38 deletions(-) diff --git a/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean b/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean index ad88a9cc37573..4ebd9831d1dd2 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ +import Mathlib.Analysis.Analytic.IteratedFDeriv import Mathlib.Analysis.Calculus.Deriv.Pow import Mathlib.Analysis.Calculus.MeanValue import Mathlib.Analysis.Calculus.ContDiff.Basic @@ -32,6 +33,11 @@ requiring that the point under consideration is accumulated by points in the int These are written using ad hoc predicates `IsSymmSndFDerivAt` and `IsSymmSndFDerivWithinAt`, which increase readability of statements in differential geometry where they show up a lot. +We also deduce statements over an arbitrary field, requiring that the function is `C^2` if the field +is `ℝ` or `ℂ`, and analytic otherwise. Formally, we assume that the function is `C^n` +with `minSmoothness 𝕜 2 ≤ n`, where `minSmoothness 𝕜 i` is `i` if `𝕜` is `ℝ` or `ℂ`, +and `ω` otherwise. + ## Implementation note For the proof, we obtain an asymptotic expansion to order two of `f (x + v + w) - f (x + v)`, by @@ -57,7 +63,7 @@ rectangle are contained in `s` by convexity. The general case follows by lineari open Asymptotics Set Filter -open scoped Topology +open scoped Topology ContDiff section General @@ -149,6 +155,23 @@ theorem IsSymmSndFDerivAt.isSymmSndFDerivWithinAt (h : IsSymmSndFDerivAt 𝕜 f simp only [← isSymmSndFDerivWithinAt_univ, ← contDiffWithinAt_univ] at h hf exact h.mono_of_mem_nhdsWithin univ_mem hf hs uniqueDiffOn_univ hx +/-- If a function is analytic within a set at a point, then its second derivative is symmetric. -/ +theorem ContDiffWithinAt.isSymmSndFDerivWithinAt_of_omega (hf : ContDiffWithinAt 𝕜 ω f s x) + (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) : + IsSymmSndFDerivWithinAt 𝕜 f s x := by + intro v w + rw [← iteratedFDerivWithin_two_apply' f hs hx, ← iteratedFDerivWithin_two_apply' f hs hx, + ← hf.iteratedFDerivWithin_comp_perm hs hx _ (Equiv.swap 0 1)] + congr + ext i + fin_cases i <;> rfl + +/-- If a function is analytic at a point, then its second derivative is symmetric. -/ +theorem ContDiffAt.isSymmSndFDerivAt_of_omega (hf : ContDiffAt 𝕜 ω f x) : + IsSymmSndFDerivAt 𝕜 f x := by + simp only [← isSymmSndFDerivWithinAt_univ, ← contDiffWithinAt_univ] at hf ⊢ + exact hf.isSymmSndFDerivWithinAt_of_omega uniqueDiffOn_univ (mem_univ _) + end General section Real @@ -407,11 +430,12 @@ end Real section IsRCLikeNormedField -variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] [IsRCLikeNormedField 𝕜] +variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E F : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] [NormedAddCommGroup F] [NormedSpace 𝕜 F] {s : Set E} {f : E → F} {x : E} -theorem second_derivative_symmetric_of_eventually {f' : E → E →L[𝕜] F} {x : E} +theorem second_derivative_symmetric_of_eventually [IsRCLikeNormedField 𝕜] + {f' : E → E →L[𝕜] F} {x : E} {f'' : E →L[𝕜] E →L[𝕜] F} (hf : ∀ᶠ y in 𝓝 x, HasFDerivAt f (f' y) y) (hx : HasFDerivAt f' f'' x) (v w : E) : f'' v w = f'' w v := by let _ := IsRCLikeNormedField.rclike 𝕜 @@ -441,37 +465,78 @@ theorem second_derivative_symmetric_of_eventually {f' : E → E →L[𝕜] F} {x /-- If a function is differentiable, and has two derivatives at `x`, then the second derivative is symmetric. -/ -theorem second_derivative_symmetric {f' : E → E →L[𝕜] F} {f'' : E →L[𝕜] E →L[𝕜] F} {x : E} +theorem second_derivative_symmetric [IsRCLikeNormedField 𝕜] + {f' : E → E →L[𝕜] F} {f'' : E →L[𝕜] E →L[𝕜] F} {x : E} (hf : ∀ y, HasFDerivAt f (f' y) y) (hx : HasFDerivAt f' f'' x) (v w : E) : f'' v w = f'' w v := second_derivative_symmetric_of_eventually (Filter.Eventually.of_forall hf) hx v w -/-- If a function is `C^2` at a point, then its second derivative there is symmetric. -/ -theorem ContDiffAt.isSymmSndFDerivAt {n : WithTop ℕ∞} (hf : ContDiffAt 𝕜 n f x) (hn : 2 ≤ n) : - IsSymmSndFDerivAt 𝕜 f x := by - intro v w - apply second_derivative_symmetric_of_eventually (f := f) (f' := fderiv 𝕜 f) (x := x) - · obtain ⟨u, hu, h'u⟩ : ∃ u ∈ 𝓝 x, ContDiffOn 𝕜 2 f u := - (hf.of_le hn).contDiffOn (m := 2) le_rfl (by simp) - rcases mem_nhds_iff.1 hu with ⟨v, vu, v_open, xv⟩ - filter_upwards [v_open.mem_nhds xv] with y hy - have : DifferentiableAt 𝕜 f y := by - have := (h'u.mono vu y hy).contDiffAt (v_open.mem_nhds hy) - exact this.differentiableAt one_le_two - exact DifferentiableAt.hasFDerivAt this - · have : DifferentiableAt 𝕜 (fderiv 𝕜 f) x := by - apply ContDiffAt.differentiableAt _ le_rfl - exact hf.fderiv_right hn - exact DifferentiableAt.hasFDerivAt this +open scoped Classical in +variable (𝕜) in +/-- `minSmoothness 𝕜 n` is the minimal smoothness exponent larger than or equal to `n` for which +one can do serious calculus in `𝕜`. If `𝕜` is `ℝ` or `ℂ`, this is just `n`. Otherwise, +this is `ω` as only analytic functions are well behaved on `ℚₚ`, say. -/ +noncomputable irreducible_def minSmoothness (n : WithTop ℕ∞) := + if IsRCLikeNormedField 𝕜 then n else ω + +@[simp] lemma minSmoothness_of_isRCLikeNormedField [h : IsRCLikeNormedField 𝕜] {n : WithTop ℕ∞} : + minSmoothness 𝕜 n = n := by + simp [minSmoothness, h] + +lemma le_minSmoothness {n : WithTop ℕ∞} : n ≤ minSmoothness 𝕜 n := by + simp only [minSmoothness] + split_ifs <;> simp + +/-- If `minSmoothness 𝕜 m ≤ n` for some (finite) integer `m`, then one can +find `n' ∈ [minSmoothness 𝕜 m, n]` which is not `∞`: over `ℝ` or `ℂ`, just take `m`, and otherwise +just take `ω`. The interest of this technical lemma is that, if a function is `C^{n'}` at a point +for `n' ≠ ∞`, then it is `C^{n'}` on a neighborhood of the point (this property fails only +in `C^∞` smoothness, see `ContDiffWithinAt.contDiffOn`). -/ +lemma exist_minSmoothness_le_ne_infty {n : WithTop ℕ∞} {m : ℕ} (hm : minSmoothness 𝕜 m ≤ n) : + ∃ n', minSmoothness 𝕜 m ≤ n' ∧ n' ≤ n ∧ n' ≠ ∞ := by + simp only [minSmoothness] at hm ⊢ + split_ifs with h + · simp only [h, ↓reduceIte] at hm + exact ⟨m, le_rfl, hm, by simp⟩ + · simp only [h, ↓reduceIte, top_le_iff] at hm + refine ⟨ω, le_rfl, by simp [hm], by simp⟩ + +/-- If a function is `C^2` at a point, then its second derivative there is symmetric. Over a field +different from `ℝ` or `ℂ`, we should require that the function is analytic. -/ +theorem ContDiffAt.isSymmSndFDerivAt {n : WithTop ℕ∞} + (hf : ContDiffAt 𝕜 n f x) (hn : minSmoothness 𝕜 2 ≤ n) : IsSymmSndFDerivAt 𝕜 f x := by + by_cases h : IsRCLikeNormedField 𝕜 + -- First deal with the `ℝ` or `ℂ` case, where `C^2` is enough. + · intro v w + apply second_derivative_symmetric_of_eventually (f := f) (f' := fderiv 𝕜 f) (x := x) + · obtain ⟨u, hu, h'u⟩ : ∃ u ∈ 𝓝 x, ContDiffOn 𝕜 2 f u := + (hf.of_le hn).contDiffOn (m := 2) le_minSmoothness (by simp) + rcases mem_nhds_iff.1 hu with ⟨v, vu, v_open, xv⟩ + filter_upwards [v_open.mem_nhds xv] with y hy + have : DifferentiableAt 𝕜 f y := by + have := (h'u.mono vu y hy).contDiffAt (v_open.mem_nhds hy) + exact this.differentiableAt one_le_two + exact DifferentiableAt.hasFDerivAt this + · have : DifferentiableAt 𝕜 (fderiv 𝕜 f) x := by + apply ContDiffAt.differentiableAt _ le_rfl + exact hf.fderiv_right (le_minSmoothness.trans hn) + exact DifferentiableAt.hasFDerivAt this + -- then deal with the case of an arbitrary field, with analytic functions. + · simp only [minSmoothness, h, ↓reduceIte, top_le_iff] at hn + apply ContDiffAt.isSymmSndFDerivAt_of_omega + simpa [hn] using hf /-- If a function is `C^2` within a set at a point, and accumulated by points in the interior -of the set, then its second derivative there is symmetric. -/ -theorem ContDiffWithinAt.isSymmSndFDerivWithinAt {n : WithTop ℕ∞} (hf : ContDiffWithinAt 𝕜 n f s x) - (hn : 2 ≤ n) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ closure (interior s)) (h'x : x ∈ s) : +of the set, then its second derivative there is symmetric. Over a field +different from `ℝ` or `ℂ`, we should require that the function is analytic. -/ +theorem ContDiffWithinAt.isSymmSndFDerivWithinAt {n : WithTop ℕ∞} + (hf : ContDiffWithinAt 𝕜 n f s x) (hn : minSmoothness 𝕜 2 ≤ n) + (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ closure (interior s)) (h'x : x ∈ s) : IsSymmSndFDerivWithinAt 𝕜 f s x := by /- We argue that, at interior points, the second derivative is symmetric, and moreover by continuity it converges to the second derivative at `x`. Therefore, the latter is also symmetric. -/ - rcases (hf.of_le hn).contDiffOn' le_rfl (by simp) with ⟨u, u_open, xu, hu⟩ + obtain ⟨m, hm, hmn, m_ne⟩ := exist_minSmoothness_le_ne_infty hn + rcases (hf.of_le hmn).contDiffOn' le_rfl (by simp [m_ne]) with ⟨u, u_open, xu, hu⟩ simp only [insert_eq_of_mem h'x] at hu have h'u : UniqueDiffOn 𝕜 (s ∩ u) := hs.inter u_open obtain ⟨y, hy, y_lim⟩ : ∃ y, (∀ (n : ℕ), y n ∈ interior s) ∧ Tendsto y atTop (𝓝 x) := @@ -483,7 +548,7 @@ theorem ContDiffWithinAt.isSymmSndFDerivWithinAt {n : WithTop ℕ∞} (hf : Cont apply mem_of_superset (isOpen_interior.mem_nhds (hy k)) exact interior_subset have : IsSymmSndFDerivAt 𝕜 f (y k) := by - apply ContDiffAt.isSymmSndFDerivAt _ le_rfl + apply ContDiffAt.isSymmSndFDerivAt _ (n := m) hm apply (hu (y k) ⟨(interior_subset (hy k)), hk⟩).contDiffAt exact inter_mem s_mem (u_open.mem_nhds hk) intro v w @@ -491,7 +556,8 @@ theorem ContDiffWithinAt.isSymmSndFDerivWithinAt {n : WithTop ℕ∞} (hf : Cont exact this v w have A : ContinuousOn (fderivWithin 𝕜 (fderivWithin 𝕜 f s) s) (s ∩ u) := by have : ContinuousOn (fderivWithin 𝕜 (fderivWithin 𝕜 f (s ∩ u)) (s ∩ u)) (s ∩ u) := - ((hu.fderivWithin h'u (m := 1) le_rfl).fderivWithin h'u (m := 0) le_rfl).continuousOn + ((hu.fderivWithin h'u (m := 1) (le_minSmoothness.trans hm)).fderivWithin h'u + (m := 0) le_rfl).continuousOn apply this.congr intro y hy apply fderivWithin_fderivWithin_eq_of_eventuallyEq diff --git a/Mathlib/Analysis/Calculus/VectorField.lean b/Mathlib/Analysis/Calculus/VectorField.lean index 0581fb044d1d1..7bd74e0209292 100644 --- a/Mathlib/Analysis/Calculus/VectorField.lean +++ b/Mathlib/Analysis/Calculus/VectorField.lean @@ -33,7 +33,7 @@ open scoped Topology noncomputable section -variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] +variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {n : WithTop ℕ∞} {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] {G : Type*} [NormedAddCommGroup G] [NormedSpace 𝕜 G] @@ -325,26 +325,29 @@ lemma leibniz_identity_lieBracketWithin_of_isSymmSndFDerivWithinAt /-- The Lie bracket of vector fields in vector spaces satisfies the Leibniz identity `[U, [V, W]] = [[U, V], W] + [V, [U, W]]`. -/ -lemma leibniz_identity_lieBracketWithin [IsRCLikeNormedField 𝕜] {U V W : E → E} {s : Set E} {x : E} +lemma leibniz_identity_lieBracketWithin (hn : minSmoothness 𝕜 2 ≤ n) + {U V W : E → E} {s : Set E} {x : E} (hs : UniqueDiffOn 𝕜 s) (h'x : x ∈ closure (interior s)) (hx : x ∈ s) - (hU : ContDiffWithinAt 𝕜 2 U s x) (hV : ContDiffWithinAt 𝕜 2 V s x) - (hW : ContDiffWithinAt 𝕜 2 W s x) : + (hU : ContDiffWithinAt 𝕜 n U s x) (hV : ContDiffWithinAt 𝕜 n V s x) + (hW : ContDiffWithinAt 𝕜 n W s x) : lieBracketWithin 𝕜 U (lieBracketWithin 𝕜 V W s) s x = lieBracketWithin 𝕜 (lieBracketWithin 𝕜 U V s) W s x + lieBracketWithin 𝕜 V (lieBracketWithin 𝕜 U W s) s x := by - apply leibniz_identity_lieBracketWithin_of_isSymmSndFDerivWithinAt hs hx hU hV hW - · exact hU.isSymmSndFDerivWithinAt le_rfl hs h'x hx - · exact hV.isSymmSndFDerivWithinAt le_rfl hs h'x hx - · exact hW.isSymmSndFDerivWithinAt le_rfl hs h'x hx + apply leibniz_identity_lieBracketWithin_of_isSymmSndFDerivWithinAt hs hx + (hU.of_le (le_minSmoothness.trans hn)) (hV.of_le (le_minSmoothness.trans hn)) + (hW.of_le (le_minSmoothness.trans hn)) + · exact hU.isSymmSndFDerivWithinAt hn hs h'x hx + · exact hV.isSymmSndFDerivWithinAt hn hs h'x hx + · exact hW.isSymmSndFDerivWithinAt hn hs h'x hx /-- The Lie bracket of vector fields in vector spaces satisfies the Leibniz identity `[U, [V, W]] = [[U, V], W] + [V, [U, W]]`. -/ -lemma leibniz_identity_lieBracket [IsRCLikeNormedField 𝕜] {U V W : E → E} {x : E} - (hU : ContDiffAt 𝕜 2 U x) (hV : ContDiffAt 𝕜 2 V x) (hW : ContDiffAt 𝕜 2 W x) : +lemma leibniz_identity_lieBracket (hn : minSmoothness 𝕜 2 ≤ n) {U V W : E → E} {x : E} + (hU : ContDiffAt 𝕜 n U x) (hV : ContDiffAt 𝕜 n V x) (hW : ContDiffAt 𝕜 n W x) : lieBracket 𝕜 U (lieBracket 𝕜 V W) x = lieBracket 𝕜 (lieBracket 𝕜 U V) W x + lieBracket 𝕜 V (lieBracket 𝕜 U W) x := by simp only [← lieBracketWithin_univ, ← contDiffWithinAt_univ] at hU hV hW ⊢ - exact leibniz_identity_lieBracketWithin uniqueDiffOn_univ (by simp) (mem_univ _) hU hV hW + exact leibniz_identity_lieBracketWithin hn uniqueDiffOn_univ (by simp) (mem_univ _) hU hV hW /-! @@ -568,4 +571,26 @@ lemma pullback_lieBracket_of_isSymmSndFDerivAt {f : E → F} {V W : F → F} {x exact pullbackWithin_lieBracketWithin_of_isSymmSndFDerivWithinAt hf h'f hV hW uniqueDiffOn_univ (mem_univ _) (mapsTo_univ _ _) +/-- The Lie bracket commutes with taking pullbacks. This requires the function to have symmetric +second derivative. Version in a complete space. One could also give a version avoiding +completeness but requiring that `f` is a local diffeo. -/ +lemma pullbackWithin_lieBracketWithin + {f : E → F} {V W : F → F} {x : E} {t : Set F} (hn : minSmoothness 𝕜 2 ≤ n) + (h'f : ContDiffWithinAt 𝕜 n f s x) + (hV : DifferentiableWithinAt 𝕜 V t (f x)) (hW : DifferentiableWithinAt 𝕜 W t (f x)) + (hu : UniqueDiffOn 𝕜 s) (hx : x ∈ s) (h'x : x ∈ closure (interior s)) (hst : MapsTo f s t) : + pullbackWithin 𝕜 f (lieBracketWithin 𝕜 V W t) s x + = lieBracketWithin 𝕜 (pullbackWithin 𝕜 f V s) (pullbackWithin 𝕜 f W s) s x := + pullbackWithin_lieBracketWithin_of_isSymmSndFDerivWithinAt + (h'f.isSymmSndFDerivWithinAt hn hu h'x hx) (h'f.of_le (le_minSmoothness.trans hn)) hV hW hu hx hst + +/-- The Lie bracket commutes with taking pullbacks. One could also give a version avoiding +completeness but requiring that `f` is a local diffeo. -/ +lemma pullback_lieBracket (hn : minSmoothness 𝕜 2 ≤ n) + {f : E → F} {V W : F → F} {x : E} (h'f : ContDiffAt 𝕜 n f x) + (hV : DifferentiableAt 𝕜 V (f x)) (hW : DifferentiableAt 𝕜 W (f x)) : + pullback 𝕜 f (lieBracket 𝕜 V W) x = lieBracket 𝕜 (pullback 𝕜 f V) (pullback 𝕜 f W) x := + pullback_lieBracket_of_isSymmSndFDerivAt (h'f.isSymmSndFDerivAt hn) + (h'f.of_le (le_minSmoothness.trans hn)) hV hW + end VectorField From 5a518c2b23c8db2280543613fff522941d39bf28 Mon Sep 17 00:00:00 2001 From: Mitchell Lee Date: Fri, 20 Dec 2024 05:45:14 +0000 Subject: [PATCH 819/829] feat: evaluate Chebyshev polynomials at small integers (#19954) We prove the following facts about Chebyshev polynomials: - $T_n(1) = 1$ - $T_n(-1) = (-1)^n$ - $U_n(1) = n + 1$ - $U_n(-1) = (-1)^n (n + 1)$ - $C_n(2) = 2$ - $C_n(-2) = 2 \cdot (-1)^n$ - $S_n(2) = n + 1$ - $S_n(-2) = (-1)^n (n + 1)$ Simplifications of the proofs are very welcome. --- Mathlib/RingTheory/Polynomial/Chebyshev.lean | 123 +++++++++++++++++++ 1 file changed, 123 insertions(+) diff --git a/Mathlib/RingTheory/Polynomial/Chebyshev.lean b/Mathlib/RingTheory/Polynomial/Chebyshev.lean index c795e148a61ce..d18e42fb4533f 100644 --- a/Mathlib/RingTheory/Polynomial/Chebyshev.lean +++ b/Mathlib/RingTheory/Polynomial/Chebyshev.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Julian Kuelshammer, Heather Macbeth, Mitchell Lee -/ import Mathlib.Algebra.Polynomial.Derivative +import Mathlib.Algebra.Ring.NegOnePow import Mathlib.Tactic.LinearCombination /-! @@ -129,6 +130,32 @@ theorem T_natAbs (n : ℤ) : T R n.natAbs = T R n := by theorem T_neg_two : T R (-2) = 2 * X ^ 2 - 1 := by simp [T_two] +@[simp] +theorem T_eval_one (n : ℤ) : (T R n).eval 1 = 1 := by + induction n using Polynomial.Chebyshev.induct with + | zero => simp + | one => simp + | add_two n ih1 ih2 => simp [T_add_two, ih1, ih2]; norm_num + | neg_add_one n ih1 ih2 => simp [T_sub_one, -T_neg, ih1, ih2]; norm_num + +@[simp] +theorem T_eval_neg_one (n : ℤ) : (T R n).eval (-1) = n.negOnePow := by + induction n using Polynomial.Chebyshev.induct with + | zero => simp + | one => simp + | add_two n ih1 ih2 => + simp only [T_add_two, eval_sub, eval_mul, eval_ofNat, eval_X, mul_neg, mul_one, ih1, + Int.negOnePow_add, Int.negOnePow_one, Units.val_neg, Int.cast_neg, neg_mul, neg_neg, ih2, + Int.negOnePow_def 2] + norm_cast + norm_num + ring + | neg_add_one n ih1 ih2 => + simp only [T_sub_one, eval_sub, eval_mul, eval_ofNat, eval_X, mul_neg, mul_one, ih1, neg_mul, + ih2, Int.negOnePow_add, Int.negOnePow_one, Units.val_neg, Int.cast_neg, sub_neg_eq_add, + Int.negOnePow_sub] + ring + /-- `U n` is the `n`-th Chebyshev polynomial of the second kind. -/ -- Well-founded definitions are now irreducible by default; -- as this was implemented before this change, @@ -194,6 +221,41 @@ theorem U_neg (n : ℤ) : U R (-n) = -U R (n - 2) := by simpa [sub_sub] using U_ theorem U_neg_sub_two (n : ℤ) : U R (-n - 2) = -U R n := by simpa [sub_eq_add_neg, add_comm] using U_neg R (n + 2) +@[simp] +theorem U_eval_one (n : ℤ) : (U R n).eval 1 = n + 1 := by + induction n using Polynomial.Chebyshev.induct with + | zero => simp + | one => simp; norm_num + | add_two n ih1 ih2 => + simp only [U_add_two, eval_sub, eval_mul, eval_ofNat, eval_X, mul_one, ih1, + Int.cast_add, Int.cast_natCast, Int.cast_one, ih2, Int.cast_ofNat] + ring + | neg_add_one n ih1 ih2 => + simp only [U_sub_one, eval_sub, eval_mul, eval_ofNat, eval_X, mul_one, + ih1, Int.cast_neg, Int.cast_natCast, ih2, Int.cast_add, Int.cast_one, Int.cast_sub, + sub_add_cancel] + ring + +@[simp] +theorem U_eval_neg_one (n : ℤ) : (U R n).eval (-1) = n.negOnePow * (n + 1) := by + induction n using Polynomial.Chebyshev.induct with + | zero => simp + | one => simp; norm_num + | add_two n ih1 ih2 => + simp only [U_add_two, eval_sub, eval_mul, eval_ofNat, eval_X, mul_neg, mul_one, ih1, + Int.cast_add, Int.cast_natCast, Int.cast_one, neg_mul, ih2, Int.cast_ofNat, Int.negOnePow_add, + Int.negOnePow_def 2] + norm_cast + norm_num + ring + | neg_add_one n ih1 ih2 => + simp only [U_sub_one, eval_sub, eval_mul, eval_ofNat, eval_X, mul_neg, mul_one, ih1, + Int.cast_neg, Int.cast_natCast, Int.negOnePow_neg, neg_mul, ih2, Int.cast_add, Int.cast_one, + Int.cast_sub, sub_add_cancel, Int.negOnePow_sub, Int.negOnePow_add] + norm_cast + norm_num + ring + theorem U_eq_X_mul_U_add_T (n : ℤ) : U R (n + 1) = X * U R n + T R (n + 1) := by induction n using Polynomial.Chebyshev.induct with | zero => simp [two_mul] @@ -289,6 +351,32 @@ theorem C_comp_two_mul_X (n : ℤ) : (C R n).comp (2 * X) = 2 * T R n := by simp_rw [C_sub_one, T_sub_one, sub_comp, mul_comp, X_comp, ih1, ih2] ring +@[simp] +theorem C_eval_two (n : ℤ) : (C R n).eval 2 = 2 := by + induction n using Polynomial.Chebyshev.induct with + | zero => simp + | one => simp + | add_two n ih1 ih2 => simp [C_add_two, ih1, ih2]; norm_num + | neg_add_one n ih1 ih2 => simp [C_sub_one, -C_neg, ih1, ih2]; norm_num + +@[simp] +theorem C_eval_neg_two (n : ℤ) : (C R n).eval (-2) = 2 * n.negOnePow := by + induction n using Polynomial.Chebyshev.induct with + | zero => simp + | one => simp + | add_two n ih1 ih2 => + simp only [C_add_two, eval_sub, eval_mul, eval_ofNat, eval_X, mul_neg, mul_one, ih1, + Int.negOnePow_add, Int.negOnePow_one, Units.val_neg, Int.cast_neg, neg_mul, neg_neg, ih2, + Int.negOnePow_def 2] + norm_cast + norm_num + ring + | neg_add_one n ih1 ih2 => + simp only [C_sub_one, eval_sub, eval_mul, eval_ofNat, eval_X, mul_neg, mul_one, ih1, neg_mul, + ih2, Int.negOnePow_add, Int.negOnePow_one, Units.val_neg, Int.cast_neg, sub_neg_eq_add, + Int.negOnePow_sub] + ring + theorem C_eq_two_mul_T_comp_half_mul_X [Invertible (2 : R)] (n : ℤ) : C R n = 2 * (T R n).comp (Polynomial.C ⅟2 * X) := by have := congr_arg (·.comp (Polynomial.C ⅟2 * X)) (C_comp_two_mul_X R n) @@ -365,6 +453,41 @@ theorem S_neg (n : ℤ) : S R (-n) = -S R (n - 2) := by simpa [sub_sub] using S_ theorem S_neg_sub_two (n : ℤ) : S R (-n - 2) = -S R n := by simpa [sub_eq_add_neg, add_comm] using S_neg R (n + 2) + @[simp] +theorem S_eval_two (n : ℤ) : (S R n).eval 2 = n + 1 := by + induction n using Polynomial.Chebyshev.induct with + | zero => simp + | one => simp; norm_num + | add_two n ih1 ih2 => + simp only [S_add_two, eval_sub, eval_mul, eval_ofNat, eval_X, mul_one, ih1, + Int.cast_add, Int.cast_natCast, Int.cast_one, ih2, Int.cast_ofNat] + ring + | neg_add_one n ih1 ih2 => + simp only [S_sub_one, eval_sub, eval_mul, eval_ofNat, eval_X, mul_one, + ih1, Int.cast_neg, Int.cast_natCast, ih2, Int.cast_add, Int.cast_one, Int.cast_sub, + sub_add_cancel] + ring + +@[simp] +theorem S_eval_neg_two (n : ℤ) : (S R n).eval (-2) = n.negOnePow * (n + 1) := by + induction n using Polynomial.Chebyshev.induct with + | zero => simp + | one => simp; norm_num + | add_two n ih1 ih2 => + simp only [S_add_two, eval_sub, eval_mul, eval_ofNat, eval_X, mul_neg, mul_one, ih1, + Int.cast_add, Int.cast_natCast, Int.cast_one, neg_mul, ih2, Int.cast_ofNat, Int.negOnePow_add, + Int.negOnePow_def 2] + norm_cast + norm_num + ring + | neg_add_one n ih1 ih2 => + simp only [S_sub_one, eval_sub, eval_mul, eval_ofNat, eval_X, mul_neg, mul_one, ih1, + Int.cast_neg, Int.cast_natCast, Int.negOnePow_neg, neg_mul, ih2, Int.cast_add, Int.cast_one, + Int.cast_sub, sub_add_cancel, Int.negOnePow_sub, Int.negOnePow_add] + norm_cast + norm_num + ring + theorem S_comp_two_mul_X (n : ℤ) : (S R n).comp (2 * X) = U R n := by induction n using Polynomial.Chebyshev.induct with | zero => simp From 1b8a2a5186cdc9bd69dc4ac266e6246625560849 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 20 Dec 2024 05:45:15 +0000 Subject: [PATCH 820/829] chore(*): drop some `[Decidable*]` assumptions (#19988) Found by the linter in #10235 --- Mathlib/Algebra/BigOperators/Group/Finset.lean | 6 +++--- Mathlib/Algebra/BigOperators/Ring.lean | 8 ++++---- Mathlib/Algebra/MvPolynomial/Degrees.lean | 4 ++-- Mathlib/CategoryTheory/Limits/Shapes/PiProd.lean | 11 ++++++----- Mathlib/LinearAlgebra/TensorProduct/Basis.lean | 14 ++++++++++++-- 5 files changed, 27 insertions(+), 16 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index ae179a61b91c4..d8de2d605f7b1 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -427,11 +427,11 @@ lemma prod_filter_not_mul_prod_filter (s : Finset α) (p : α → Prop) [Decidab rw [mul_comm, prod_filter_mul_prod_filter_not] @[to_additive] -theorem prod_filter_xor (p q : α → Prop) [DecidableEq α] [DecidablePred p] [DecidablePred q] : +theorem prod_filter_xor (p q : α → Prop) [DecidablePred p] [DecidablePred q] : (∏ x ∈ s with (Xor' (p x) (q x)), f x) = (∏ x ∈ s with (p x ∧ ¬ q x), f x) * (∏ x ∈ s with (q x ∧ ¬ p x), f x) := by - rw [← prod_union (disjoint_filter_and_not_filter _ _), ← filter_or] - rfl + classical rw [← prod_union (disjoint_filter_and_not_filter _ _), ← filter_or] + simp only [Xor'] section ToList diff --git a/Mathlib/Algebra/BigOperators/Ring.lean b/Mathlib/Algebra/BigOperators/Ring.lean index 6e86cf047cd5e..a548b362058b9 100644 --- a/Mathlib/Algebra/BigOperators/Ring.lean +++ b/Mathlib/Algebra/BigOperators/Ring.lean @@ -178,15 +178,15 @@ theorem prod_add (f g : ι → α) (s : Finset ι) : simp only [mem_filter, mem_sdiff, not_and, not_exists, and_congr_right_iff] tauto) +end DecidableEq + theorem prod_one_add {f : ι → α} (s : Finset ι) : ∏ i ∈ s, (1 + f i) = ∑ t ∈ s.powerset, ∏ i ∈ t, f i := by - simp only [add_comm (1 : α), prod_add, prod_const_one, mul_one] + classical simp only [add_comm (1 : α), prod_add, prod_const_one, mul_one] theorem prod_add_one {f : ι → α} (s : Finset ι) : ∏ i ∈ s, (f i + 1) = ∑ t ∈ s.powerset, ∏ i ∈ t, f i := by - simp only [prod_add, prod_const_one, mul_one] - -end DecidableEq + classical simp only [prod_add, prod_const_one, mul_one] /-- `∏ i, (f i + g i) = (∏ i, f i) + ∑ i, g i * (∏ j < i, f j + g j) * (∏ j > i, f j)`. -/ theorem prod_add_ordered [LinearOrder ι] (s : Finset ι) (f g : ι → α) : diff --git a/Mathlib/Algebra/MvPolynomial/Degrees.lean b/Mathlib/Algebra/MvPolynomial/Degrees.lean index 5227f6ddbbe60..55b5f391e4aff 100644 --- a/Mathlib/Algebra/MvPolynomial/Degrees.lean +++ b/Mathlib/Algebra/MvPolynomial/Degrees.lean @@ -253,9 +253,9 @@ theorem monomial_le_degreeOf (i : σ) {f : MvPolynomial σ R} {m : σ →₀ ℕ rw [degreeOf_eq_sup i] apply Finset.le_sup h_m -lemma degreeOf_monomial_eq [DecidableEq σ] (s : σ →₀ ℕ) (i : σ) {a : R} (ha : a ≠ 0) : +lemma degreeOf_monomial_eq (s : σ →₀ ℕ) (i : σ) {a : R} (ha : a ≠ 0) : (monomial s a).degreeOf i = s i := by - rw [degreeOf_def, degrees_monomial_eq _ _ ha, Finsupp.count_toMultiset] + classical rw [degreeOf_def, degrees_monomial_eq _ _ ha, Finsupp.count_toMultiset] -- TODO we can prove equality with `NoZeroDivisors R` theorem degreeOf_mul_le (i : σ) (f g : MvPolynomial σ R) : diff --git a/Mathlib/CategoryTheory/Limits/Shapes/PiProd.lean b/Mathlib/CategoryTheory/Limits/Shapes/PiProd.lean index 6027232e5928d..03b61c5cb2cd0 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/PiProd.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/PiProd.lean @@ -17,7 +17,7 @@ and its complement. namespace CategoryTheory.Limits variable {C I : Type*} [Category C] {X Y : I → C} - (f : (i : I) → X i ⟶ Y i) (P : I → Prop) [∀ i, Decidable (P i)] + (f : (i : I) → X i ⟶ Y i) (P : I → Prop) [HasProduct X] [HasProduct Y] [HasProduct (fun (i : {x : I // P x}) ↦ X i.val)] [HasProduct (fun (i : {x : I // ¬ P x}) ↦ X i.val)] @@ -39,7 +39,8 @@ variable (X) in A product indexed by `I` is a binary product of the products indexed by a subset of `I` and its complement. -/ -noncomputable def Pi.binaryFanOfPropIsLimit : IsLimit (Pi.binaryFanOfProp X P) := +noncomputable def Pi.binaryFanOfPropIsLimit [∀ i, Decidable (P i)] : + IsLimit (Pi.binaryFanOfProp X P) := BinaryFan.isLimitMk (fun s ↦ Pi.lift fun b ↦ if h : P b then s.π.app ⟨WalkingPair.left⟩ ≫ Pi.π (fun (i : {x : I // P x}) ↦ X i.val) ⟨b, h⟩ else @@ -51,12 +52,12 @@ noncomputable def Pi.binaryFanOfPropIsLimit : IsLimit (Pi.binaryFanOfProp X P) : · simp [← h₂, dif_neg h]) lemma hasBinaryProduct_of_products : HasBinaryProduct (∏ᶜ (fun (i : {x : I // P x}) ↦ X i.val)) - (∏ᶜ (fun (i : {x : I // ¬ P x}) ↦ X i.val)) := - ⟨Pi.binaryFanOfProp X P, Pi.binaryFanOfPropIsLimit X P⟩ + (∏ᶜ (fun (i : {x : I // ¬ P x}) ↦ X i.val)) := by + classical exact ⟨Pi.binaryFanOfProp X P, Pi.binaryFanOfPropIsLimit X P⟩ attribute [local instance] hasBinaryProduct_of_products -lemma Pi.map_eq_prod_map : Pi.map f = +lemma Pi.map_eq_prod_map [∀ i, Decidable (P i)] : Pi.map f = ((Pi.binaryFanOfPropIsLimit X P).conePointUniqueUpToIso (prodIsProd _ _)).hom ≫ prod.map (Pi.map (fun (i : {x : I // P x}) ↦ f i.val)) (Pi.map (fun (i : {x : I // ¬ P x}) ↦ f i.val)) ≫ diff --git a/Mathlib/LinearAlgebra/TensorProduct/Basis.lean b/Mathlib/LinearAlgebra/TensorProduct/Basis.lean index 00a58ae8b3df6..f4794bfbc014f 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Basis.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Basis.lean @@ -100,13 +100,17 @@ lemma TensorProduct.equivFinsuppOfBasisRight_symm_apply (b : κ →₀ M) : (TensorProduct.equivFinsuppOfBasisRight 𝒞).symm b = b.sum fun i m ↦ m ⊗ₜ 𝒞 i := congr($(TensorProduct.equivFinsuppOfBasisRight_symm 𝒞) b) +omit [DecidableEq κ] in lemma TensorProduct.sum_tmul_basis_right_injective : Function.Injective (Finsupp.lsum R fun i ↦ (TensorProduct.mk R M N).flip (𝒞 i)) := + have := Classical.decEq κ (equivFinsuppOfBasisRight_symm (M := M) 𝒞).symm ▸ (TensorProduct.equivFinsuppOfBasisRight 𝒞).symm.injective +omit [DecidableEq κ] in lemma TensorProduct.sum_tmul_basis_right_eq_zero (b : κ →₀ M) (h : (b.sum fun i m ↦ m ⊗ₜ[R] 𝒞 i) = 0) : b = 0 := + have := Classical.decEq κ (TensorProduct.equivFinsuppOfBasisRight 𝒞).symm.injective (a₂ := 0) <| by simpa /-- @@ -138,26 +142,32 @@ lemma TensorProduct.equivFinsuppOfBasisLeft_symm_apply (b : ι →₀ N) : (TensorProduct.equivFinsuppOfBasisLeft ℬ).symm b = b.sum fun i n ↦ ℬ i ⊗ₜ n := congr($(TensorProduct.equivFinsuppOfBasisLeft_symm ℬ) b) +omit [DecidableEq κ] in /-- Elements in `M ⊗ N` can be represented by sum of elements in `M` tensor elements of basis of `N`. -/ lemma TensorProduct.eq_repr_basis_right : ∃ b : κ →₀ M, b.sum (fun i m ↦ m ⊗ₜ 𝒞 i) = x := by - simpa using (TensorProduct.equivFinsuppOfBasisRight 𝒞).symm.surjective x + classical simpa using (TensorProduct.equivFinsuppOfBasisRight 𝒞).symm.surjective x +omit [DecidableEq ι] in /-- Elements in `M ⊗ N` can be represented by sum of elements of basis of `M` tensor elements of `N`.-/ lemma TensorProduct.eq_repr_basis_left : ∃ (c : ι →₀ N), (c.sum fun i n ↦ ℬ i ⊗ₜ n) = x := by - obtain ⟨c, rfl⟩ := (TensorProduct.equivFinsuppOfBasisLeft ℬ).symm.surjective x + classical obtain ⟨c, rfl⟩ := (TensorProduct.equivFinsuppOfBasisLeft ℬ).symm.surjective x exact ⟨c, (TensorProduct.comm R M N).injective <| by simp [Finsupp.sum]⟩ +omit [DecidableEq ι] in lemma TensorProduct.sum_tmul_basis_left_injective : Function.Injective (Finsupp.lsum R fun i ↦ (TensorProduct.mk R M N) (ℬ i)) := + have := Classical.decEq ι (equivFinsuppOfBasisLeft_symm (N := N) ℬ).symm ▸ (TensorProduct.equivFinsuppOfBasisLeft ℬ).symm.injective +omit [DecidableEq ι] in lemma TensorProduct.sum_tmul_basis_left_eq_zero (b : ι →₀ N) (h : (b.sum fun i n ↦ ℬ i ⊗ₜ[R] n) = 0) : b = 0 := + have := Classical.decEq ι (TensorProduct.equivFinsuppOfBasisLeft ℬ).symm.injective (a₂ := 0) <| by simpa end From af8d133b3172e865899d7c37a7df09c805f22130 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Fri, 20 Dec 2024 05:45:17 +0000 Subject: [PATCH 821/829] chore: generalize MulEquiv.prime_iff (#20093) Also put the simpler form on the RHS and simplify the proofs that use it. --- Mathlib/Algebra/Prime/Lemmas.lean | 8 +++--- Mathlib/RingTheory/Polynomial/Basic.lean | 26 +++++++------------ .../UniqueFactorizationDomain/Basic.lean | 2 +- 3 files changed, 16 insertions(+), 20 deletions(-) diff --git a/Mathlib/Algebra/Prime/Lemmas.lean b/Mathlib/Algebra/Prime/Lemmas.lean index 1df2748e62898..6461b52529e62 100644 --- a/Mathlib/Algebra/Prime/Lemmas.lean +++ b/Mathlib/Algebra/Prime/Lemmas.lean @@ -53,9 +53,11 @@ theorem comap_prime (hinv : ∀ a, g (f a : N) = a) (hp : Prime (f p)) : Prime p · intro h convert ← map_dvd g h <;> apply hinv⟩ -theorem MulEquiv.prime_iff (e : M ≃* N) : Prime p ↔ Prime (e p) := - ⟨fun h => (comap_prime e.symm e fun a => by simp) <| (e.symm_apply_apply p).substr h, - comap_prime e e.symm fun a => by simp⟩ +theorem MulEquiv.prime_iff {E : Type*} [EquivLike E M N] [MulEquivClass E M N] (e : E) : + Prime (e p) ↔ Prime p := by + let e := MulEquivClass.toMulEquiv e + exact ⟨comap_prime e e.symm fun a => by simp, + fun h => (comap_prime e.symm e fun a => by simp) <| (e.symm_apply_apply p).substr h⟩ end Map diff --git a/Mathlib/RingTheory/Polynomial/Basic.lean b/Mathlib/RingTheory/Polynomial/Basic.lean index 5513563c4b0da..3a13bb5216d8a 100644 --- a/Mathlib/RingTheory/Polynomial/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Basic.lean @@ -810,18 +810,15 @@ namespace MvPolynomial private theorem prime_C_iff_of_fintype {R : Type u} (σ : Type v) {r : R} [CommRing R] [Fintype σ] : Prime (C r : MvPolynomial σ R) ↔ Prime r := by - rw [(renameEquiv R (Fintype.equivFin σ)).toMulEquiv.prime_iff] + rw [← MulEquiv.prime_iff (renameEquiv R (Fintype.equivFin σ))] convert_to Prime (C r) ↔ _ · congr! - apply rename_C - · symm - induction' Fintype.card σ with d hd - · exact (isEmptyAlgEquiv R (Fin 0)).toMulEquiv.symm.prime_iff - · rw [hd, ← Polynomial.prime_C_iff] - convert (finSuccEquiv R d).toMulEquiv.symm.prime_iff (p := Polynomial.C (C r)) - rw [← finSuccEquiv_comp_C_eq_C] - simp_rw [RingHom.coe_comp, RingHom.coe_coe, Function.comp_apply, MulEquiv.symm_mk, - AlgEquiv.toEquiv_eq_coe, AlgEquiv.symm_toEquiv_eq_symm, MulEquiv.coe_mk, EquivLike.coe_coe] + simp only [renameEquiv_apply, algHom_C, algebraMap_eq] + · induction' Fintype.card σ with d hd + · exact MulEquiv.prime_iff (isEmptyAlgEquiv R (Fin 0)).symm (p := r) + · convert MulEquiv.prime_iff (finSuccEquiv R d).symm (p := Polynomial.C (C r)) + · simp [← finSuccEquiv_comp_C_eq_C] + · simp [← hd, Polynomial.prime_C_iff] theorem prime_C_iff : Prime (C r : MvPolynomial σ R) ↔ Prime r := ⟨comap_prime C constantCoeff (constantCoeff_C _), fun hr => @@ -864,12 +861,9 @@ theorem prime_rename_iff (s : Set σ) {p : MvPolynomial s R} : iterToSum_C_X, renameEquiv_apply, Equiv.coe_trans, Equiv.sumComm_apply, Sum.swap_inr, Equiv.Set.sumCompl_apply_inl] apply_fun (· p) at this - simp_rw [AlgHom.toRingHom_eq_coe, RingHom.coe_coe] at this - rw [← prime_C_iff, eqv.toMulEquiv.prime_iff, this] - simp only [MulEquiv.coe_mk, AlgEquiv.toEquiv_eq_coe, EquivLike.coe_coe, AlgEquiv.trans_apply, - MvPolynomial.sumAlgEquiv_symm_apply, renameEquiv_apply, Equiv.coe_trans, Equiv.sumComm_apply, - AlgEquiv.toAlgHom_eq_coe, AlgEquiv.toAlgHom_toRingHom, RingHom.coe_comp, RingHom.coe_coe, - AlgEquiv.coe_trans, Function.comp_apply] + simp only [AlgHom.toRingHom_eq_coe, RingHom.coe_coe, AlgEquiv.toAlgHom_eq_coe, + AlgEquiv.toAlgHom_toRingHom, RingHom.coe_comp, Function.comp_apply] at this + rw [this, MulEquiv.prime_iff, prime_C_iff] end MvPolynomial diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Basic.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Basic.lean index 9d04cd5c5c1f9..72776c6b23c75 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain/Basic.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Basic.lean @@ -391,7 +391,7 @@ theorem MulEquiv.uniqueFactorizationMonoid (e : α ≃* β) (hα : UniqueFactori exact ⟨w.map e, fun b hb => let ⟨c, hc, he⟩ := Multiset.mem_map.1 hb - he ▸ e.prime_iff.1 (hp c hc), + he ▸ e.prime_iff.2 (hp c hc), Units.map e.toMonoidHom u, by rw [Multiset.prod_hom, toMonoidHom_eq_coe, Units.coe_map, MonoidHom.coe_coe, ← map_mul e, h, From aceed3da123350f8c4643676ff7ce06def0030b2 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Fri, 20 Dec 2024 06:52:40 +0000 Subject: [PATCH 822/829] chore: move normed order homs to Analysis (#20003) --- Mathlib.lean | 4 ++-- .../Hom/Normed.lean => Analysis/Normed/Order/Hom/Basic.lean} | 2 +- Mathlib/{Algebra => Analysis/Normed}/Order/Hom/Ultra.lean | 2 +- Mathlib/Data/Real/IsNonarchimedean.lean | 2 +- 4 files changed, 5 insertions(+), 5 deletions(-) rename Mathlib/{Algebra/Order/Hom/Normed.lean => Analysis/Normed/Order/Hom/Basic.lean} (97%) rename Mathlib/{Algebra => Analysis/Normed}/Order/Hom/Ultra.lean (97%) diff --git a/Mathlib.lean b/Mathlib.lean index a43ecaafda619..f85f8bf378a99 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -697,9 +697,7 @@ import Mathlib.Algebra.Order.GroupWithZero.Unbundled.Lemmas import Mathlib.Algebra.Order.GroupWithZero.WithZero import Mathlib.Algebra.Order.Hom.Basic import Mathlib.Algebra.Order.Hom.Monoid -import Mathlib.Algebra.Order.Hom.Normed import Mathlib.Algebra.Order.Hom.Ring -import Mathlib.Algebra.Order.Hom.Ultra import Mathlib.Algebra.Order.Interval.Basic import Mathlib.Algebra.Order.Interval.Finset import Mathlib.Algebra.Order.Interval.Multiset @@ -1446,6 +1444,8 @@ import Mathlib.Analysis.Normed.Operator.Compact import Mathlib.Analysis.Normed.Operator.ContinuousLinearMap import Mathlib.Analysis.Normed.Operator.LinearIsometry import Mathlib.Analysis.Normed.Order.Basic +import Mathlib.Analysis.Normed.Order.Hom.Basic +import Mathlib.Analysis.Normed.Order.Hom.Ultra import Mathlib.Analysis.Normed.Order.Lattice import Mathlib.Analysis.Normed.Order.UpperLower import Mathlib.Analysis.Normed.Ring.IsPowMulFaithful diff --git a/Mathlib/Algebra/Order/Hom/Normed.lean b/Mathlib/Analysis/Normed/Order/Hom/Basic.lean similarity index 97% rename from Mathlib/Algebra/Order/Hom/Normed.lean rename to Mathlib/Analysis/Normed/Order/Hom/Basic.lean index b44ee621b367a..b7aa436182858 100644 --- a/Mathlib/Algebra/Order/Hom/Normed.lean +++ b/Mathlib/Analysis/Normed/Order/Hom/Basic.lean @@ -12,7 +12,7 @@ import Mathlib.Analysis.Normed.Group.Basic This file defines constructions that upgrade `(Comm)Group` to `(Semi)Normed(Comm)Group` using a `Group(Semi)normClass` when the codomain is the reals. -See `Mathlib.Algebra.Order.Hom.Ultra` for further upgrades to nonarchimedean normed groups. +See `Mathlib.Analysis.Normed.Order.Hom.Ultra` for further upgrades to nonarchimedean normed groups. -/ diff --git a/Mathlib/Algebra/Order/Hom/Ultra.lean b/Mathlib/Analysis/Normed/Order/Hom/Ultra.lean similarity index 97% rename from Mathlib/Algebra/Order/Hom/Ultra.lean rename to Mathlib/Analysis/Normed/Order/Hom/Ultra.lean index a0a03cd82334a..b28160bad854a 100644 --- a/Mathlib/Algebra/Order/Hom/Ultra.lean +++ b/Mathlib/Analysis/Normed/Order/Hom/Ultra.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Yakov Pechersky. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yakov Pechersky -/ -import Mathlib.Algebra.Order.Hom.Normed +import Mathlib.Analysis.Normed.Order.Hom.Basic import Mathlib.Topology.MetricSpace.Ultra.Basic /-! diff --git a/Mathlib/Data/Real/IsNonarchimedean.lean b/Mathlib/Data/Real/IsNonarchimedean.lean index 6b8017ac1d629..47bfc1f9722cc 100644 --- a/Mathlib/Data/Real/IsNonarchimedean.lean +++ b/Mathlib/Data/Real/IsNonarchimedean.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 María Inés de Frutos-Fernández. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: María Inés de Frutos-Fernández -/ -import Mathlib.Algebra.Order.Hom.Ultra +import Mathlib.Analysis.Normed.Order.Hom.Ultra import Mathlib.Analysis.Normed.Group.Ultra import Mathlib.Data.Nat.Choose.Sum From 8854cbcbb3fbf505347e332c8709e5b0230eb9a3 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 20 Dec 2024 06:52:42 +0000 Subject: [PATCH 823/829] chore: further backports for leanprover/lean4#6397 (#20098) Co-authored-by: Johan Commelin --- Mathlib/Algebra/FreeMonoid/Count.lean | 16 +++++++++++++--- Mathlib/Algebra/Group/Defs.lean | 2 +- .../Gaussian/GaussianIntegral.lean | 2 +- Mathlib/CategoryTheory/Functor/OfSequence.lean | 2 +- Mathlib/Combinatorics/Quiver/SingleObj.lean | 2 +- Mathlib/Computability/PartrecCode.lean | 4 ++-- Mathlib/Computability/RegularExpressions.lean | 2 +- Mathlib/Computability/TMToPartrec.lean | 8 ++++---- Mathlib/Data/List/Basic.lean | 4 ++-- Mathlib/Data/List/DropRight.lean | 5 +++-- Mathlib/Data/List/EditDistance/Defs.lean | 4 ++-- Mathlib/Data/List/Indexes.lean | 6 +++--- Mathlib/Data/List/Range.lean | 5 +++-- Mathlib/Data/Nat/Multiplicity.lean | 2 +- Mathlib/Data/Num/Lemmas.lean | 2 +- Mathlib/Data/Ordmap/Ordset.lean | 2 +- Mathlib/Data/TypeVec.lean | 6 +++--- Mathlib/Data/W/Constructions.lean | 4 ++-- Mathlib/Deprecated/LazyList.lean | 14 +++++++------- Mathlib/GroupTheory/CoprodI.lean | 4 ++-- Mathlib/GroupTheory/Coxeter/Inversion.lean | 2 +- .../ModelTheory/Algebra/Field/IsAlgClosed.lean | 7 +++---- Mathlib/ModelTheory/Semantics.lean | 12 ++++++------ Mathlib/RingTheory/FractionalIdeal/Basic.lean | 2 +- Mathlib/SetTheory/ZFC/Basic.lean | 3 +++ Mathlib/SetTheory/ZFC/Rank.lean | 2 +- .../Topology/Category/Profinite/Nobeling.lean | 7 +++---- 27 files changed, 72 insertions(+), 59 deletions(-) diff --git a/Mathlib/Algebra/FreeMonoid/Count.lean b/Mathlib/Algebra/FreeMonoid/Count.lean index 53b2c0e4c51c3..d9add268027ba 100644 --- a/Mathlib/Algebra/FreeMonoid/Count.lean +++ b/Mathlib/Algebra/FreeMonoid/Count.lean @@ -14,6 +14,15 @@ In this file we define `FreeMonoid.countP`, `FreeMonoid.count`, `FreeAddMonoid.c additive homomorphisms from `FreeMonoid` and `FreeAddMonoid`. We do not use `to_additive` because it can't map `Multiplicative ℕ` to `ℕ`. + +## TODO + +There is lots of defeq abuse here of `FreeAddMonoid α = List α`, e.g. in statements like +``` +theorem countP_apply (l : FreeAddMonoid α) : countP p l = List.countP p l := rfl +``` +This needs cleaning up. + -/ variable {α : Type*} (p : α → Prop) [DecidablePred p] @@ -27,7 +36,8 @@ def countP : FreeAddMonoid α →+ ℕ where map_add' := List.countP_append _ theorem countP_of (x : α) : countP p (of x) = if p x = true then 1 else 0 := by - simp [countP, List.countP, List.countP.go] + change List.countP p [x] = _ + simp [List.countP_cons] theorem countP_apply (l : FreeAddMonoid α) : countP p l = List.countP p l := rfl @@ -36,8 +46,8 @@ theorem countP_apply (l : FreeAddMonoid α) : countP p l = List.countP p l := rf def count [DecidableEq α] (x : α) : FreeAddMonoid α →+ ℕ := countP (· = x) theorem count_of [DecidableEq α] (x y : α) : count x (of y) = (Pi.single x 1 : α → ℕ) y := by - simp [Pi.single, Function.update, count, countP, List.countP, List.countP.go, - Bool.beq_eq_decide_eq] + change List.count x [y] = _ + simp [Pi.single, Function.update, List.count_cons] theorem count_apply [DecidableEq α] (x : α) (l : FreeAddMonoid α) : count x l = List.count x l := rfl diff --git a/Mathlib/Algebra/Group/Defs.lean b/Mathlib/Algebra/Group/Defs.lean index 394c147b9bf64..bc799629080db 100644 --- a/Mathlib/Algebra/Group/Defs.lean +++ b/Mathlib/Algebra/Group/Defs.lean @@ -451,7 +451,7 @@ theorem npowRec'_two_mul {M : Type*} [Semigroup M] [One M] (k : ℕ) (m : M) : match k' with | 0 => rfl | 1 => simp [npowRec'] - | k + 2 => simp [npowRec', ← mul_assoc, ← ih] + | k + 2 => simp [npowRec', ← mul_assoc, Nat.mul_add, ← ih] @[to_additive] theorem npowRec'_mul_comm {M : Type*} [Semigroup M] [One M] {k : ℕ} (k0 : k ≠ 0) (m : M) : diff --git a/Mathlib/Analysis/SpecialFunctions/Gaussian/GaussianIntegral.lean b/Mathlib/Analysis/SpecialFunctions/Gaussian/GaussianIntegral.lean index 812fe616e1ee9..6124cdfeef08b 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gaussian/GaussianIntegral.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gaussian/GaussianIntegral.lean @@ -361,7 +361,7 @@ lemma Real.Gamma_nat_add_one_add_half (k : ℕ) : induction k with | zero => simp [-one_div, add_comm (1 : ℝ), Gamma_add_one, Gamma_one_half_eq]; ring | succ k ih => - rw [add_right_comm, Gamma_add_one (by positivity), Nat.cast_add, Nat.cast_one, ih] + rw [add_right_comm, Gamma_add_one (by positivity), Nat.cast_add, Nat.cast_one, ih, Nat.mul_add] field_simp ring diff --git a/Mathlib/CategoryTheory/Functor/OfSequence.lean b/Mathlib/CategoryTheory/Functor/OfSequence.lean index c1f5788cec33b..e1bf6334adeac 100644 --- a/Mathlib/CategoryTheory/Functor/OfSequence.lean +++ b/Mathlib/CategoryTheory/Functor/OfSequence.lean @@ -81,7 +81,7 @@ lemma map_comp (i j k : ℕ) (hij : i ≤ j) (hjk : j ≤ k) : · omega · obtain rfl : j = 0 := by omega rw [map_id, comp_id] - · dsimp [map] + · simp only [map, Nat.reduceAdd] rw [hj (fun n ↦ f (n + 1)) (k + 1) (by omega) (by omega)] obtain _|j := j all_goals simp [map] diff --git a/Mathlib/Combinatorics/Quiver/SingleObj.lean b/Mathlib/Combinatorics/Quiver/SingleObj.lean index 0f8ecb66a9de8..d87aa9c958488 100644 --- a/Mathlib/Combinatorics/Quiver/SingleObj.lean +++ b/Mathlib/Combinatorics/Quiver/SingleObj.lean @@ -113,7 +113,7 @@ theorem listToPath_pathToList {x : SingleObj α} (p : Path (star α) x) : listToPath (pathToList p) = p.cast rfl ext := by induction p with | nil => rfl - | cons _ _ ih => dsimp at *; rw [ih] + | cons _ _ ih => dsimp [pathToList] at *; rw [ih] theorem pathToList_listToPath (l : List α) : pathToList (listToPath l) = l := by induction l with diff --git a/Mathlib/Computability/PartrecCode.lean b/Mathlib/Computability/PartrecCode.lean index 52df16d8893d0..f86ef6f073202 100644 --- a/Mathlib/Computability/PartrecCode.lean +++ b/Mathlib/Computability/PartrecCode.lean @@ -495,10 +495,10 @@ theorem eval_const : ∀ n m, eval (Code.const n) m = Part.some n | n + 1, m => by simp! [eval_const n m] @[simp] -theorem eval_id (n) : eval Code.id n = Part.some n := by simp! [Seq.seq] +theorem eval_id (n) : eval Code.id n = Part.some n := by simp! [Seq.seq, Code.id] @[simp] -theorem eval_curry (c n x) : eval (curry c n) x = eval c (Nat.pair n x) := by simp! [Seq.seq] +theorem eval_curry (c n x) : eval (curry c n) x = eval c (Nat.pair n x) := by simp! [Seq.seq, curry] theorem const_prim : Primrec Code.const := (_root_.Primrec.id.nat_iterate (_root_.Primrec.const zero) diff --git a/Mathlib/Computability/RegularExpressions.lean b/Mathlib/Computability/RegularExpressions.lean index da53a917cd11d..b816d00833eee 100644 --- a/Mathlib/Computability/RegularExpressions.lean +++ b/Mathlib/Computability/RegularExpressions.lean @@ -364,7 +364,7 @@ def map (f : α → β) : RegularExpression α → RegularExpression β @[simp] protected theorem map_pow (f : α → β) (P : RegularExpression α) : ∀ n : ℕ, map f (P ^ n) = map f P ^ n - | 0 => by dsimp; rfl + | 0 => by unfold map; rfl | n + 1 => (congr_arg (· * map f P) (RegularExpression.map_pow f P n) : _) #adaptation_note /-- around nightly-2024-02-25, diff --git a/Mathlib/Computability/TMToPartrec.lean b/Mathlib/Computability/TMToPartrec.lean index 8b095162d02d9..0b08e16e8fbb1 100644 --- a/Mathlib/Computability/TMToPartrec.lean +++ b/Mathlib/Computability/TMToPartrec.lean @@ -515,11 +515,11 @@ def Cont.then : Cont → Cont → Cont theorem Cont.then_eval {k k' : Cont} {v} : (k.then k').eval v = k.eval v >>= k'.eval := by induction k generalizing v with | halt => simp only [Cont.eval, Cont.then, pure_bind] - | cons₁ => simp only [Cont.eval, bind_assoc, *] - | cons₂ => simp only [Cont.eval, *] - | comp _ _ k_ih => simp only [Cont.eval, bind_assoc, ← k_ih] + | cons₁ => simp only [Cont.eval, Cont.then, bind_assoc, *] + | cons₂ => simp only [Cont.eval, Cont.then, *] + | comp _ _ k_ih => simp only [Cont.eval, Cont.then, bind_assoc, ← k_ih] | fix _ _ k_ih => - simp only [Cont.eval, *] + simp only [Cont.eval, Cont.then, *] split_ifs <;> [rfl; simp only [← k_ih, bind_assoc]] /-- The `then k` function is a "configuration homomorphism". Its operation on states is to append diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 9c5ae0bef9266..31ba4a24600dc 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -1233,11 +1233,11 @@ theorem getElem_succ_scanl {i : ℕ} (h : i + 1 < (scanl f b l).length) : induction i generalizing b l with | zero => cases l - · simp only [length, zero_eq, lt_self_iff_false] at h + · simp only [scanl, length, zero_eq, lt_self_iff_false] at h · simp | succ i hi => cases l - · simp only [length] at h + · simp only [scanl, length] at h exact absurd h (by omega) · simp_rw [scanl_cons] rw [getElem_append_right] diff --git a/Mathlib/Data/List/DropRight.lean b/Mathlib/Data/List/DropRight.lean index ad7f1c36acb62..4de39f926b613 100644 --- a/Mathlib/Data/List/DropRight.lean +++ b/Mathlib/Data/List/DropRight.lean @@ -188,8 +188,9 @@ theorem rtakeWhile_eq_nil_iff : rtakeWhile p l = [] ↔ ∀ hl : l ≠ [], ¬p ( induction' l using List.reverseRecOn with l a · simp only [rtakeWhile, takeWhile, reverse_nil, true_iff] intro f; contradiction - · simp only [rtakeWhile, reverse_append, takeWhile, ne_eq, not_false_eq_true, - getLast_append_of_ne_nil, getLast_singleton, reduceCtorEq] + · simp only [rtakeWhile, reverse_append, reverse_cons, reverse_nil, nil_append, singleton_append, + takeWhile, ne_eq, cons_ne_self, not_false_eq_true, getLast_append_of_ne_nil, + getLast_singleton] refine ⟨fun h => ?_ , fun h => ?_⟩ · split at h <;> simp_all · simp [h] diff --git a/Mathlib/Data/List/EditDistance/Defs.lean b/Mathlib/Data/List/EditDistance/Defs.lean index a52cf793066d8..e60d98190de9e 100644 --- a/Mathlib/Data/List/EditDistance/Defs.lean +++ b/Mathlib/Data/List/EditDistance/Defs.lean @@ -267,9 +267,9 @@ theorem suffixLevenshtein_eq_tails_map (xs ys) : (suffixLevenshtein C xs ys).1 = xs.tails.map fun xs' => levenshtein C xs' ys := by induction xs with | nil => - simp only [List.map, suffixLevenshtein_nil'] + simp only [suffixLevenshtein_nil', List.tails, List.map_cons, List.map] | cons x xs ih => - simp only [List.map, suffixLevenshtein_cons₁, ih] + simp only [suffixLevenshtein_cons₁, ih, List.tails, List.map_cons] @[simp] theorem levenshtein_nil_nil : levenshtein C [] [] = 0 := by diff --git a/Mathlib/Data/List/Indexes.lean b/Mathlib/Data/List/Indexes.lean index dfa2e8e7d9973..e9680556a8109 100644 --- a/Mathlib/Data/List/Indexes.lean +++ b/Mathlib/Data/List/Indexes.lean @@ -55,8 +55,8 @@ theorem map_enumFrom_eq_zipWith : ∀ (l : List α) (n : ℕ) (f : ℕ → α rw [this]; rfl · cases' l with head tail · contradiction - · simp only [map, uncurry_apply_pair, range_succ_eq_map, zipWith, Nat.zero_add, - zipWith_map_left] + · simp only [enumFrom_cons, map_cons, range_succ_eq_map, zipWith_cons_cons, + Nat.zero_add, zipWith_map_left, true_and] rw [ih] · suffices (fun i ↦ f (i + (n + 1))) = ((fun i ↦ f (i + n)) ∘ Nat.succ) by rw [this] @@ -253,7 +253,7 @@ theorem mapIdxMGo_eq_mapIdxMAuxSpec cases as · rfl · contradiction - simp only [this, mapIdxM.go, mapIdxMAuxSpec, List.traverse, map_pure, append_nil] + simp only [this, mapIdxM.go, mapIdxMAuxSpec, enumFrom_nil, List.traverse, map_pure, append_nil] · match as with | nil => contradiction | cons head tail => diff --git a/Mathlib/Data/List/Range.lean b/Mathlib/Data/List/Range.lean index e133f0b507d6d..9f633f23c4b44 100644 --- a/Mathlib/Data/List/Range.lean +++ b/Mathlib/Data/List/Range.lean @@ -82,7 +82,7 @@ theorem ranges_length (l : List ℕ) : induction l with | nil => simp only [ranges, map_nil] | cons a l hl => -- (a :: l) - simp only [map, length_range, map_map, cons.injEq, true_and] + simp only [ranges, map_cons, length_range, map_map, cons.injEq, true_and] conv_rhs => rw [← hl] apply map_congr_left intro s _ @@ -93,7 +93,8 @@ set_option linter.deprecated false in @[deprecated "Use `List.ranges_flatten`." (since := "2024-10-17")] lemma ranges_flatten' : ∀ l : List ℕ, l.ranges.flatten = range (Nat.sum l) | [] => rfl - | a :: l => by simp only [Nat.sum_cons, flatten, ← map_flatten, ranges_flatten', range_add] + | a :: l => by + simp only [ranges, flatten_cons, ← map_flatten, ranges_flatten', Nat.sum_cons, range_add] @[deprecated (since := "2024-10-15")] alias ranges_join' := ranges_flatten' diff --git a/Mathlib/Data/Nat/Multiplicity.lean b/Mathlib/Data/Nat/Multiplicity.lean index ece38a314c41c..99e32e92a2a92 100644 --- a/Mathlib/Data/Nat/Multiplicity.lean +++ b/Mathlib/Data/Nat/Multiplicity.lean @@ -278,7 +278,7 @@ theorem emultiplicity_two_factorial_lt : ∀ {n : ℕ} (_ : n ≠ 0), emultiplic by_cases hn : n = 0 · subst hn simp only [ne_eq, bit_eq_zero_iff, true_and, Bool.not_eq_false] at h - simp only [h, factorial, mul_one, Nat.isUnit_iff, cast_one] + simp only [bit, h, cond_true, mul_zero, zero_add, factorial_one] rw [Prime.emultiplicity_one] · exact zero_lt_one · decide diff --git a/Mathlib/Data/Num/Lemmas.lean b/Mathlib/Data/Num/Lemmas.lean index c7e98484c0e1c..777ee93d18575 100644 --- a/Mathlib/Data/Num/Lemmas.lean +++ b/Mathlib/Data/Num/Lemmas.lean @@ -861,7 +861,7 @@ theorem castNum_testBit (m n) : testBit m n = Nat.testBit m n := by | pos m => rw [cast_pos] induction' n with n IH generalizing m <;> cases' m with m m - <;> dsimp only [PosNum.testBit] + <;> simp only [PosNum.testBit] · 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] diff --git a/Mathlib/Data/Ordmap/Ordset.lean b/Mathlib/Data/Ordmap/Ordset.lean index e1d9a4f8ef4b9..5b497cb5df020 100644 --- a/Mathlib/Data/Ordmap/Ordset.lean +++ b/Mathlib/Data/Ordmap/Ordset.lean @@ -1378,7 +1378,7 @@ theorem Valid'.map_aux {β} [Preorder β] {f : α → β} (f_strict_mono : Stric simp only [map, size_nil, and_true]; apply valid'_nil cases a₁; · trivial cases a₂; · trivial - simp only [Bounded] + simp only [Option.map, Bounded] exact f_strict_mono h.ord | node _ _ _ _ t_ih_l t_ih_r => have t_ih_l' := t_ih_l h.left diff --git a/Mathlib/Data/TypeVec.lean b/Mathlib/Data/TypeVec.lean index 5c337fc5f0cd1..88be36a0d3c00 100644 --- a/Mathlib/Data/TypeVec.lean +++ b/Mathlib/Data/TypeVec.lean @@ -662,14 +662,14 @@ theorem subtypeVal_diagSub {α : TypeVec n} : subtypeVal (repeatEq α) ⊚ diagS theorem toSubtype_of_subtype {α : TypeVec n} (p : α ⟹ «repeat» n Prop) : toSubtype p ⊚ ofSubtype p = id := by ext i x - induction i <;> dsimp only [id, toSubtype, comp, ofSubtype] at * + induction i <;> simp only [id, toSubtype, comp, ofSubtype] at * simp [*] @[simp] theorem subtypeVal_toSubtype {α : TypeVec n} (p : α ⟹ «repeat» n Prop) : subtypeVal p ⊚ toSubtype p = fun _ => Subtype.val := by ext i x - induction i <;> dsimp only [toSubtype, comp, subtypeVal] at * + induction i <;> simp only [toSubtype, comp, subtypeVal] at * simp [*] @[simp] @@ -689,7 +689,7 @@ theorem toSubtype'_of_subtype' {α : TypeVec n} (r : α ⊗ α ⟹ «repeat» n theorem subtypeVal_toSubtype' {α : TypeVec n} (r : α ⊗ α ⟹ «repeat» n Prop) : subtypeVal r ⊚ toSubtype' r = fun i x => prod.mk i x.1.fst x.1.snd := by ext i x - induction i <;> dsimp only [id, toSubtype', comp, subtypeVal, prod.mk] at * + induction i <;> simp only [id, toSubtype', comp, subtypeVal, prod.mk] at * simp [*] end TypeVec diff --git a/Mathlib/Data/W/Constructions.lean b/Mathlib/Data/W/Constructions.lean index 8134ef7814e89..940f86a200b28 100644 --- a/Mathlib/Data/W/Constructions.lean +++ b/Mathlib/Data/W/Constructions.lean @@ -150,12 +150,12 @@ theorem leftInverse_list : Function.LeftInverse (ofList γ) (toList _) ext x cases x | WType.mk (Listα.cons x) f => by - simp only [ofList, leftInverse_list (f PUnit.unit), mk.injEq, heq_eq_eq, true_and] + simp only [toList, ofList, leftInverse_list (f PUnit.unit), mk.injEq, heq_eq_eq, true_and] rfl theorem rightInverse_list : Function.RightInverse (ofList γ) (toList _) | List.nil => rfl - | List.cons hd tl => by simp only [toList, rightInverse_list tl] + | List.cons hd tl => by simp [rightInverse_list tl] /-- Lists are equivalent to their associated `WType` -/ def equivList : WType (Listβ γ) ≃ List γ where diff --git a/Mathlib/Deprecated/LazyList.lean b/Mathlib/Deprecated/LazyList.lean index 5ba81d38999a8..04f7bd5f25dc5 100644 --- a/Mathlib/Deprecated/LazyList.lean +++ b/Mathlib/Deprecated/LazyList.lean @@ -50,17 +50,17 @@ instance : LawfulTraversable LazyList := by · induction xs using LazyList.rec with | nil => simp only [Functor.map, LazyList.traverse, pure, Equiv.map, listEquivLazyList, - Equiv.coe_fn_symm_mk, toList, Equiv.coe_fn_mk, ofList] + Equiv.coe_fn_symm_mk, toList, List.map_nil, Equiv.coe_fn_mk, ofList] | cons => - simpa only [Equiv.map, Functor.map, listEquivLazyList, Equiv.coe_fn_symm_mk, Equiv.coe_fn_mk, - LazyList.traverse, Seq.seq, toList, ofList, cons.injEq, true_and] + simpa only [Functor.map, LazyList.traverse, Seq.seq, Equiv.map, listEquivLazyList, + Equiv.coe_fn_symm_mk, toList, List.map_cons, Equiv.coe_fn_mk, ofList, cons.injEq, true_and] | mk _ ih => ext; apply ih - · simp only [Equiv.map, listEquivLazyList, Equiv.coe_fn_symm_mk, Equiv.coe_fn_mk, comp, - Functor.mapConst] + · simp only [Functor.mapConst, comp, Equiv.map, listEquivLazyList, Equiv.coe_fn_symm_mk, + List.map_eq_map, List.map_const, Equiv.coe_fn_mk] induction xs using LazyList.rec with - | nil => simp only [LazyList.traverse, pure, Functor.map, toList, ofList] + | nil => simp [LazyList.traverse, pure, Functor.map, toList, ofList] | cons => - simpa only [toList, ofList, LazyList.traverse, Seq.seq, Functor.map, cons.injEq, true_and] + simpa [toList, ofList, LazyList.traverse, Seq.seq, Functor.map, cons.injEq, true_and] | mk _ ih => congr; apply ih · simp only [traverse, Equiv.traverse, listEquivLazyList, Equiv.coe_fn_mk, Equiv.coe_fn_symm_mk] induction xs using LazyList.rec with diff --git a/Mathlib/GroupTheory/CoprodI.lean b/Mathlib/GroupTheory/CoprodI.lean index 5e020fad4e35f..e5bc86accb345 100644 --- a/Mathlib/GroupTheory/CoprodI.lean +++ b/Mathlib/GroupTheory/CoprodI.lean @@ -754,7 +754,7 @@ theorem replaceHead_head {i j : ι} (x : M i) (hnotone : x ≠ 1) (w : NeWord M (replaceHead x hnotone w).head = x := by induction w · rfl - · simp [*] + · simp [*, replaceHead] /-- One can multiply an element from the left to a non-empty reduced word if it does not cancel with the first element in the word. -/ @@ -766,7 +766,7 @@ theorem mulHead_head {i j : ι} (w : NeWord M i j) (x : M i) (hnotone : x * w.he (mulHead w x hnotone).head = x * w.head := by induction w · rfl - · simp [*] + · simp [*, mulHead] @[simp] theorem mulHead_prod {i j : ι} (w : NeWord M i j) (x : M i) (hnotone : x * w.head ≠ 1) : diff --git a/Mathlib/GroupTheory/Coxeter/Inversion.lean b/Mathlib/GroupTheory/Coxeter/Inversion.lean index 4603c629ae971..38f2db8c97c26 100644 --- a/Mathlib/GroupTheory/Coxeter/Inversion.lean +++ b/Mathlib/GroupTheory/Coxeter/Inversion.lean @@ -209,7 +209,7 @@ theorem rightInvSeq_concat (ω : List B) (i : B) : ris (ω.concat i) = (List.map (MulAut.conj (s i)) (ris ω)).concat (s i) := by induction' ω with j ω ih · simp - · dsimp [rightInvSeq] + · dsimp [rightInvSeq, concat] rw [ih] simp only [concat_eq_append, wordProd_append, wordProd_cons, wordProd_nil, mul_one, mul_inv_rev, inv_simple, cons_append, cons.injEq, and_true] diff --git a/Mathlib/ModelTheory/Algebra/Field/IsAlgClosed.lean b/Mathlib/ModelTheory/Algebra/Field/IsAlgClosed.lean index 54708abe18b61..cef615ebaa210 100644 --- a/Mathlib/ModelTheory/Algebra/Field/IsAlgClosed.lean +++ b/Mathlib/ModelTheory/Algebra/Field/IsAlgClosed.lean @@ -208,10 +208,9 @@ theorem finite_ACF_prime_not_realize_of_ACF_zero_realize have := modelField_of_modelACF q K let _ := compatibleRingOfModelField K have := charP_of_model_fieldOfChar q K - simp only [eqZero, Term.equal, Term.relabel, BoundedFormula.realize_not, - BoundedFormula.realize_bdEqual, Term.realize_relabel, Sum.elim_comp_inl, - realize_termOfFreeCommRing, map_natCast, Term.realize_func, CompatibleRing.funMap_zero, - ne_eq, ← CharP.charP_iff_prime_eq_zero hp] + simp only [eqZero, Term.equal, BoundedFormula.realize_not, BoundedFormula.realize_bdEqual, + Term.realize_relabel, Sum.elim_comp_inl, realize_termOfFreeCommRing, map_natCast, + realize_zero, ← CharP.charP_iff_prime_eq_zero hp] intro _ exact hqp <| CharP.eq K inferInstance inferInstance let s : Finset Nat.Primes := T0.attach.biUnion (fun φ => f φ.1 (hT0 φ.2)) diff --git a/Mathlib/ModelTheory/Semantics.lean b/Mathlib/ModelTheory/Semantics.lean index 75203851ec43e..cfb1e54ddc657 100644 --- a/Mathlib/ModelTheory/Semantics.lean +++ b/Mathlib/ModelTheory/Semantics.lean @@ -154,13 +154,13 @@ theorem realize_constantsToVars [L[[α]].Structure M] [(lhomWithConstants L α). · simp · cases n · cases f - · simp only [realize, ih, constantsOn, constantsOnFunc] + · simp only [realize, ih, constantsOn, constantsOnFunc, constantsToVars] -- Porting note: below lemma does not work with simp for some reason rw [withConstants_funMap_sum_inl] · simp only [realize, constantsToVars, Sum.elim_inl, funMap_eq_coe_constants] rfl · cases' f with _ f - · simp only [realize, ih, constantsOn, constantsOnFunc] + · simp only [realize, ih, constantsOn, constantsOnFunc, constantsToVars] -- Porting note: below lemma does not work with simp for some reason rw [withConstants_funMap_sum_inl] · exact isEmptyElim f @@ -174,7 +174,7 @@ theorem realize_varsToConstants [L[[α]].Structure M] [(lhomWithConstants L α). -- Porting note: both cases were `simp [Language.con]` · simp [Language.con, realize, funMap_eq_coe_constants] · simp [realize, constantMap] - · simp only [realize, constantsOn, constantsOnFunc, ih] + · simp only [realize, constantsOn, constantsOnFunc, ih, varsToConstants] -- Porting note: below lemma does not work with simp for some reason rw [withConstants_funMap_sum_inl] @@ -417,16 +417,16 @@ theorem realize_restrictFreeVar [DecidableEq α] {n : ℕ} {φ : L.BoundedFormul induction φ with | falsum => rfl | equal => - simp only [Realize, freeVarFinset.eq_2] + simp only [Realize, freeVarFinset.eq_2, restrictFreeVar] rw [Set.inclusion_comp_inclusion, Set.inclusion_comp_inclusion] simp | rel => - simp only [Realize, freeVarFinset.eq_3, Finset.biUnion_val] + simp only [Realize, freeVarFinset.eq_3, Finset.biUnion_val, restrictFreeVar] congr! erw [Set.inclusion_comp_inclusion _ h] simp | imp _ _ ih1 ih2 => - simp only [Realize, freeVarFinset.eq_4] + simp only [Realize, freeVarFinset.eq_4, restrictFreeVar] rw [Set.inclusion_comp_inclusion, Set.inclusion_comp_inclusion] simp [ih1, ih2] | all _ ih3 => simp [restrictFreeVar, Realize, ih3] diff --git a/Mathlib/RingTheory/FractionalIdeal/Basic.lean b/Mathlib/RingTheory/FractionalIdeal/Basic.lean index ab635bda253dd..e151dbde1d83b 100644 --- a/Mathlib/RingTheory/FractionalIdeal/Basic.lean +++ b/Mathlib/RingTheory/FractionalIdeal/Basic.lean @@ -526,7 +526,7 @@ theorem mul_eq_mul (I J : FractionalIdeal S P) : mul I J = I * J := rfl theorem mul_def (I J : FractionalIdeal S P) : - I * J = ⟨I * J, I.isFractional.mul J.isFractional⟩ := by simp only [← mul_eq_mul, mul] + I * J = ⟨I * J, I.isFractional.mul J.isFractional⟩ := by simp only [← mul_eq_mul, mul_def'] @[simp, norm_cast] theorem coe_mul (I J : FractionalIdeal S P) : (↑(I * J) : Submodule R P) = I * J := by diff --git a/Mathlib/SetTheory/ZFC/Basic.lean b/Mathlib/SetTheory/ZFC/Basic.lean index c590a930ecd31..30c195428de2f 100644 --- a/Mathlib/SetTheory/ZFC/Basic.lean +++ b/Mathlib/SetTheory/ZFC/Basic.lean @@ -308,6 +308,9 @@ instance : Inhabited PSet := instance : IsEmpty («Type» ∅) := ⟨PEmpty.elim⟩ +theorem empty_def : (∅ : PSet) = ⟨_, PEmpty.elim⟩ := by + simp [EmptyCollection.emptyCollection, PSet.empty] + @[simp] theorem not_mem_empty (x : PSet.{u}) : x ∉ (∅ : PSet.{u}) := IsEmpty.exists_iff.1 diff --git a/Mathlib/SetTheory/ZFC/Rank.lean b/Mathlib/SetTheory/ZFC/Rank.lean index 0f0e5aeafd4a9..08715c0e83eb0 100644 --- a/Mathlib/SetTheory/ZFC/Rank.lean +++ b/Mathlib/SetTheory/ZFC/Rank.lean @@ -65,7 +65,7 @@ variable {x y : PSet.{u}} rank_le_iff.2 fun _ h₁ => rank_lt_of_mem (mem_of_subset h h₁) @[simp] -theorem rank_empty : rank ∅ = 0 := by simp [rank] +theorem rank_empty : rank ∅ = 0 := by simp [empty_def, rank] @[simp] theorem rank_insert (x y : PSet) : rank (insert x y) = max (succ (rank x)) (rank y) := by diff --git a/Mathlib/Topology/Category/Profinite/Nobeling.lean b/Mathlib/Topology/Category/Profinite/Nobeling.lean index 494a50eb97186..2af0c9f25bf29 100644 --- a/Mathlib/Topology/Category/Profinite/Nobeling.lean +++ b/Mathlib/Topology/Category/Profinite/Nobeling.lean @@ -791,8 +791,7 @@ instance {α : Type*} [TopologicalSpace α] [Nonempty α] : Nontrivial (LocallyC theorem Products.isGood_nil {I} [LinearOrder I] : Products.isGood ({fun _ ↦ false} : Set (I → Bool)) Products.nil := by intro h - simp only [Products.lt_nil_empty, Products.eval, List.map, List.prod_nil, Set.image_empty, - Submodule.span_empty, Submodule.mem_bot, one_ne_zero] at h + simp [Products.eval, Products.nil] at h theorem Products.span_nil_eq_top {I} [LinearOrder I] : Submodule.span ℤ (eval ({fun _ ↦ false} : Set (I → Bool)) '' {nil}) = ⊤ := by @@ -800,7 +799,7 @@ theorem Products.span_nil_eq_top {I} [LinearOrder I] : intro f _ rw [Submodule.mem_span_singleton] refine ⟨f default, ?_⟩ - simp only [eval, List.map, List.prod_nil, zsmul_eq_mul, mul_one] + simp only [eval, List.map, List.prod_nil, zsmul_eq_mul, mul_one, Products.nil] ext x obtain rfl : x = default := by simp only [Set.default_coe_singleton, eq_iff_true_of_subsingleton] rfl @@ -835,7 +834,7 @@ instance (α : Type*) [TopologicalSpace α] : NoZeroSMulDivisors ℤ (LocallyCon theorem GoodProducts.linearIndependentSingleton {I} [LinearOrder I] : LinearIndependent ℤ (eval ({fun _ ↦ false} : Set (I → Bool))) := by refine linearIndependent_unique (eval ({fun _ ↦ false} : Set (I → Bool))) ?_ - simp only [eval, Products.eval, List.map, List.prod_nil, ne_eq, one_ne_zero, not_false_eq_true] + simp [eval, Products.eval, Products.nil, default] end Zero From e0cf109cd53a9eab08403037e1190790ed020a9d Mon Sep 17 00:00:00 2001 From: FMLJohn <417963616@qq.com> Date: Fri, 20 Dec 2024 08:07:02 +0000 Subject: [PATCH 824/829] feat(RingTheory/Polynomial/HilbertPoly): `Polynomial.hilbertPoly_linearMap` (#19897) --- .../RingTheory/Polynomial/HilbertPoly.lean | 45 +++++++++++++++---- 1 file changed, 37 insertions(+), 8 deletions(-) diff --git a/Mathlib/RingTheory/Polynomial/HilbertPoly.lean b/Mathlib/RingTheory/Polynomial/HilbertPoly.lean index 32146950375d5..21a5c87a28f37 100644 --- a/Mathlib/RingTheory/Polynomial/HilbertPoly.lean +++ b/Mathlib/RingTheory/Polynomial/HilbertPoly.lean @@ -19,7 +19,7 @@ given any `p : F[X]` and `d : ℕ`, there exists some `h : F[X]` such that for a `n : ℕ`, `h(n)` is equal to the coefficient of `Xⁿ` in the power series expansion of `p/(1 - X)ᵈ`. This `h` is unique and is denoted as `Polynomial.hilbertPoly p d`. -For example, given `d : ℕ`, the power series expansion of `1/(1-X)ᵈ⁺¹` in `F[X]` +For example, given `d : ℕ`, the power series expansion of `1/(1 - X)ᵈ⁺¹` in `F[X]` is `Σₙ ((d + n).choose d)Xⁿ`, which equals `Σₙ ((n + 1)···(n + d)/d!)Xⁿ` and hence `Polynomial.hilbertPoly (1 : F[X]) (d + 1)` is the polynomial `(X + 1)···(X + d)/d!`. Note that if `d! = 0` in `F`, then the polynomial `(X + 1)···(X + d)/d!` no longer works, so we do not want @@ -106,22 +106,51 @@ noncomputable def hilbertPoly (p : F[X]) : (d : ℕ) → F[X] | 0 => 0 | d + 1 => ∑ i in p.support, (p.coeff i) • preHilbertPoly F d i -variable (F) in -lemma hilbertPoly_zero_nat (d : ℕ) : hilbertPoly (0 : F[X]) d = 0 := by +lemma hilbertPoly_zero_left (d : ℕ) : hilbertPoly (0 : F[X]) d = 0 := by delta hilbertPoly; induction d with | zero => simp only | succ d _ => simp only [coeff_zero, zero_smul, Finset.sum_const_zero] -lemma hilbertPoly_poly_zero (p : F[X]) : hilbertPoly p 0 = 0 := rfl +lemma hilbertPoly_zero_right (p : F[X]) : hilbertPoly p 0 = 0 := rfl -lemma hilbertPoly_poly_succ (p : F[X]) (d : ℕ) : +lemma hilbertPoly_succ (p : F[X]) (d : ℕ) : hilbertPoly p (d + 1) = ∑ i in p.support, (p.coeff i) • preHilbertPoly F d i := rfl -variable (F) in lemma hilbertPoly_X_pow_succ (d k : ℕ) : hilbertPoly ((X : F[X]) ^ k) (d + 1) = preHilbertPoly F d k := by delta hilbertPoly; simp +lemma hilbertPoly_add_left (p q : F[X]) (d : ℕ) : + hilbertPoly (p + q) d = hilbertPoly p d + hilbertPoly q d := by + delta hilbertPoly + induction d with + | zero => simp only [add_zero] + | succ d _ => + simp only [← coeff_add] + rw [← sum_def _ fun _ r => r • _] + exact sum_add_index _ _ _ (fun _ => zero_smul ..) (fun _ _ _ => add_smul ..) + +lemma hilbertPoly_smul (a : F) (p : F[X]) (d : ℕ) : + hilbertPoly (a • p) d = a • hilbertPoly p d := by + delta hilbertPoly + induction d with + | zero => simp only [smul_zero] + | succ d _ => + simp only + rw [← sum_def _ fun _ r => r • _, ← sum_def _ fun _ r => r • _, Polynomial.smul_sum, + sum_smul_index' _ _ _ fun i => zero_smul F (preHilbertPoly F d i)] + simp only [smul_assoc] + +variable (F) in +/-- +The function that sends any `p : F[X]` to `Polynomial.hilbertPoly p d` is an `F`-linear map from +`F[X]` to `F[X]`. +-/ +noncomputable def hilbertPoly_linearMap (d : ℕ) : F[X] →ₗ[F] F[X] where + toFun p := hilbertPoly p d + map_add' p q := hilbertPoly_add_left p q d + map_smul' r p := hilbertPoly_smul r p d + variable [CharZero F] /-- @@ -203,7 +232,7 @@ lemma hilbertPoly_eq_zero_of_le_rootMultiplicity_one {p : F[X]} {d : ℕ} (hdp : d ≤ p.rootMultiplicity 1) : hilbertPoly p d = 0 := by by_cases hp : p = 0 - · rw [hp, hilbertPoly_zero_nat] + · rw [hp, hilbertPoly_zero_left] · rcases exists_eq_pow_rootMultiplicity_mul_and_not_dvd p hp 1 with ⟨q, hq1, hq2⟩ have heq : p = q * (- 1) ^ p.rootMultiplicity 1 * (1 - X) ^ p.rootMultiplicity 1 := by simp only [mul_assoc, ← mul_pow, neg_mul, one_mul, neg_sub] @@ -238,7 +267,7 @@ theorem natDegree_hilbertPoly_of_ne_zero have hp : p ≠ 0 := by intro h rw [h] at hh - exact hh (hilbertPoly_zero_nat F d) + exact hh (hilbertPoly_zero_left d) have hpd : p.rootMultiplicity 1 < d := by by_contra h exact hh (hilbertPoly_eq_zero_of_le_rootMultiplicity_one <| not_lt.1 h) From 77833b207e9bd0bec7270a66a1034081f12afe6f Mon Sep 17 00:00:00 2001 From: Seewoo Lee Date: Fri, 20 Dec 2024 08:07:04 +0000 Subject: [PATCH 825/829] feat(Counterexamples): irrational a, b with rational a^b (#19931) Co-authored-by: Seewoo Lee <49933279+seewoo5@users.noreply.github.com> --- Counterexamples.lean | 1 + .../IrrationalPowerOfIrrational.lean | 37 +++++++++++++++++++ 2 files changed, 38 insertions(+) create mode 100644 Counterexamples/IrrationalPowerOfIrrational.lean diff --git a/Counterexamples.lean b/Counterexamples.lean index d6c0c3af5e0a5..c0cefc537076c 100644 --- a/Counterexamples.lean +++ b/Counterexamples.lean @@ -7,6 +7,7 @@ import Counterexamples.DiscreteTopologyNonDiscreteUniformity import Counterexamples.GameMultiplication import Counterexamples.Girard import Counterexamples.HomogeneousPrimeNotPrime +import Counterexamples.IrrationalPowerOfIrrational import Counterexamples.LinearOrderWithPosMulPosEqZero import Counterexamples.MapFloor import Counterexamples.MonicNonRegular diff --git a/Counterexamples/IrrationalPowerOfIrrational.lean b/Counterexamples/IrrationalPowerOfIrrational.lean new file mode 100644 index 0000000000000..75e822d68b149 --- /dev/null +++ b/Counterexamples/IrrationalPowerOfIrrational.lean @@ -0,0 +1,37 @@ +/- +Copyright (c) 2024 Seewoo Lee. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Seewoo Lee +-/ +import Mathlib.Data.Real.Irrational +import Mathlib.Analysis.SpecialFunctions.Pow.NNReal + +/- +# Irrational power of irrational numbers are not necessarily irrational. + +Prove that there exist irrational numbers `a`, `b` such that `a^b` is rational. + +We use the following famous argument (based on the law of excluded middle and irrationality of √2). +Consider `c = √2^√2`. If `c` is rational, we are done. +If `c` is irrational, then `c^√2 = 2` is rational, so we are done. +-/ + + +open Classical Real + +namespace Counterexample + +/-- +There exist irrational `a`, `b` with rational `a^b`. +Note that the positivity assumption on `a` is imposed because of the definition of `rpow` for +negative bases. See `Real.rpow_def_of_neg` for more details. +-/ +theorem not_irrational_rpow : + ¬ ∀ a b : ℝ, Irrational a → Irrational b → 0 < a → Irrational (a ^ b) := by + push_neg + by_cases hc : Irrational (√2 ^ √2) + · use (√2 ^ √2), √2, hc, irrational_sqrt_two, by positivity + rw [← rpow_mul, mul_self_sqrt, rpow_two, sq_sqrt] <;> norm_num + · use √2, √2, irrational_sqrt_two, irrational_sqrt_two, by positivity, hc + +end Counterexample From f3381fd64d26c20321b5a7404c9c789798550bfe Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Fri, 20 Dec 2024 08:07:05 +0000 Subject: [PATCH 826/829] chore(Norm/Transitivity): replace `RingHom.mapMatrix` by `Matrix.map` (#20009) because `Matrix.map` is the simple normal form --- Mathlib/RingTheory/Norm/Transitivity.lean | 39 ++++++++++++----------- 1 file changed, 20 insertions(+), 19 deletions(-) diff --git a/Mathlib/RingTheory/Norm/Transitivity.lean b/Mathlib/RingTheory/Norm/Transitivity.lean index 7476a9d8ce06c..1ef8e6e741b67 100644 --- a/Mathlib/RingTheory/Norm/Transitivity.lean +++ b/Mathlib/RingTheory/Norm/Transitivity.lean @@ -75,7 +75,7 @@ lemma mul_auxMat_toSquareBlock_eq : set_option quotPrecheck false in /-- The upper-left block of `M * aux M k`. -/ -notation "mulAuxMatBlock" => (M * auxMat M k).toSquareBlock (· = k) False +scoped notation "mulAuxMatBlock" => (M * auxMat M k).toSquareBlock (· = k) False lemma det_mul_corner_pow : M.det * M k k ^ (Fintype.card m - 1) = M k k * (mulAuxMatBlock).det := by @@ -91,10 +91,11 @@ lemma det_mul_corner_pow : /-- A matrix with X added to the corner. -/ noncomputable def cornerAddX : Matrix m m S[X] := - (diagonal fun i ↦ if i = k then X else 0) + C.mapMatrix M + (diagonal fun i ↦ if i = k then X else 0) + M.map C variable [Fintype n] (f : S →+* Matrix n n R) +omit [Fintype m] in lemma polyToMatrix_cornerAddX : f.polyToMatrix (cornerAddX M k k k) = (-f (M k k)).charmatrix := by simp [cornerAddX, Matrix.add_apply, charmatrix, @@ -106,35 +107,35 @@ lemma eval_zero_det_det : eval 0 (f.polyToMatrix (cornerAddX M k).det).det = (f congr; ext; simp [cornerAddX, diagonal, apply_ite] lemma eval_zero_comp_det : - eval 0 (comp m m n n R[X] <| f.polyToMatrix.mapMatrix <| cornerAddX M k).det = - (comp m m n n R <| f.mapMatrix M).det := by + eval 0 (comp m m n n R[X] <| (cornerAddX M k).map f.polyToMatrix).det = + (comp m m n n R <| M.map f).det := by simp_rw [← coe_evalRingHom, RingHom.map_det, ← compRingEquiv_apply, ← RingEquiv.coe_toRingHom, - ← RingHom.comp_apply, ← RingHom.comp_assoc, evalRingHom_mapMatrix_comp_compRingEquiv, - RingHom.comp_assoc, RingHom.mapMatrix_comp, evalRingHom_mapMatrix_comp_polyToMatrix, - ← RingHom.mapMatrix_comp, RingHom.comp_apply] + ← RingHom.mapMatrix_apply, ← RingHom.comp_apply, ← RingHom.comp_assoc, + evalRingHom_mapMatrix_comp_compRingEquiv, RingHom.comp_assoc, RingHom.mapMatrix_comp, + evalRingHom_mapMatrix_comp_polyToMatrix, ← RingHom.mapMatrix_comp, RingHom.comp_apply] congr with i j simp [cornerAddX, diagonal, apply_ite] theorem comp_det_mul_pow : - ((f.mapMatrix M).comp m m n n R).det * (f (M k k)).det ^ (Fintype.card m - 1) = - (f (M k k)).det * ((f.mapMatrix mulAuxMatBlock).comp _ _ n n R).det := by - trans ((f.mapMatrix (M * auxMat M k)).compRingEquiv m n R).det - · simp_rw [_root_.map_mul, det_mul, f.mapMatrix_apply, compRingEquiv_apply, - ((auxMat_blockTriangular M k).map f).comp.det_fintype, Fintype.prod_Prop, + ((M.map f).comp m m n n R).det * (f (M k k)).det ^ (Fintype.card m - 1) = + (f (M k k)).det * (((mulAuxMatBlock).map f).comp _ _ n n R).det := by + trans (((M * auxMat M k).map f).comp m m n n R).det + · simp_rw [← f.mapMatrix_apply, ← compRingEquiv_apply, _root_.map_mul, det_mul, f.mapMatrix_apply, + compRingEquiv_apply, ((auxMat_blockTriangular M k).map f).comp.det_fintype, Fintype.prod_Prop, comp_toSquareBlock (b := (· ≠ k)), det_reindex_self, map_toSquareBlock, auxMat_toSquareBlock_eq, auxMat_toSquareBlock_ne, smul_one_eq_diagonal, ← diagonal_one, diagonal_map (map_zero _), comp_diagonal, det_reindex_self] simp - · simp_rw [f.mapMatrix_apply, compRingEquiv_apply, ((mul_auxMat_blockTriangular M k).map f).comp - |>.det_fintype, Fintype.prod_Prop, comp_toSquareBlock (b := (· = k)), det_reindex_self, - map_toSquareBlock, mul_auxMat_toSquareBlock_eq, smul_one_eq_diagonal, + · simp_rw [((mul_auxMat_blockTriangular M k).map f).comp.det_fintype, Fintype.prod_Prop, + comp_toSquareBlock (b := (· = k)), det_reindex_self, map_toSquareBlock, + mul_auxMat_toSquareBlock_eq, smul_one_eq_diagonal, diagonal_map (map_zero _), comp_diagonal, det_reindex_self] simp variable {M f} in lemma det_det_aux - (ih : ∀ M, (f M.det).det = ((f.mapMatrix M).comp {a // (a = k) = False} _ n n R).det) : - ((f M.det).det - ((f.mapMatrix M).comp m m n n R).det) * + (ih : ∀ M, (f (det M)).det = ((M.map f).comp {a // (a = k) = False} _ n n R).det) : + ((f M.det).det - ((M.map f).comp m m n n R).det) * (f (M k k)).det ^ (Fintype.card m - 1) = 0 := by rw [sub_mul, comp_det_mul_pow, ← det_pow, ← map_pow, ← det_mul, ← _root_.map_mul, det_mul_corner_pow, _root_.map_mul, det_mul, ih, sub_self] @@ -148,7 +149,7 @@ a block matrix with commuting, equal-sized, square blocks can be computed by tak twice in a row: first take the determinant over the commutative ring generated by the blocks (`S` here), then take the determinant over the base ring. -/ theorem Matrix.det_det [Fintype m] [Fintype n] (f : S →+* Matrix n n R) : - (f M.det).det = ((f.mapMatrix M).comp m m n n R).det := by + (f M.det).det = ((M.map f).comp m m n n R).det := by set l := Fintype.card m with hl clear_value l; revert R S m induction' l with l ih <;> intro R S m _ _ M _ _ f card @@ -157,7 +158,7 @@ theorem Matrix.det_det [Fintype m] [Fintype n] (f : S →+* Matrix n n R) : have ⟨k⟩ := Fintype.card_pos_iff.mp (l.succ_pos.trans_eq card) let f' := f.polyToMatrix let M' := cornerAddX M k - have : (f' M'.det).det = ((f'.mapMatrix M').comp m m n n R[X]).det := by + have : (f' M'.det).det = ((M'.map f').comp m m n n R[X]).det := by refine sub_eq_zero.mp <| mem_nonZeroDivisors_iff.mp (pow_mem ?_ _) _ (det_det_aux k fun M ↦ ih _ _ <| by simp [← card]) rw [polyToMatrix_cornerAddX, ← charpoly] From faa9c3fafe11017615f1fdc4812b03ffc581a565 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Fri, 20 Dec 2024 08:18:05 +0000 Subject: [PATCH 827/829] chore: move `RepresentationTheory/Action` to `CategoryTheory/Action` (#20002) Actions are not linear enough for representation theory. This commit moves the `RepresentationTheory/Action` directory to `CategoryTheory/Action` and adjusts the imports accordingly. This also means that the theory of Galois categories no longer needs to import from `RepresentationTheory/`. --- Mathlib.lean | 10 +++++----- Mathlib/Algebra/Category/MonCat/Basic.lean | 4 ++-- .../Action/Basic.lean | 0 .../Action/Concrete.lean | 2 +- .../Action/Continuous.lean | 2 +- .../Action/Limits.lean | 2 +- .../Action/Monoidal.lean | 4 ++-- Mathlib/CategoryTheory/Galois/Examples.lean | 4 ++-- Mathlib/RepresentationTheory/Rep.lean | 2 +- 9 files changed, 15 insertions(+), 15 deletions(-) rename Mathlib/{RepresentationTheory => CategoryTheory}/Action/Basic.lean (100%) rename Mathlib/{RepresentationTheory => CategoryTheory}/Action/Concrete.lean (99%) rename Mathlib/{RepresentationTheory => CategoryTheory}/Action/Continuous.lean (98%) rename Mathlib/{RepresentationTheory => CategoryTheory}/Action/Limits.lean (99%) rename Mathlib/{RepresentationTheory => CategoryTheory}/Action/Monoidal.lean (99%) diff --git a/Mathlib.lean b/Mathlib.lean index f85f8bf378a99..e9cefd4b02346 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1597,6 +1597,11 @@ import Mathlib.CategoryTheory.Abelian.RightDerived import Mathlib.CategoryTheory.Abelian.Subobject import Mathlib.CategoryTheory.Abelian.Transfer import Mathlib.CategoryTheory.Action +import Mathlib.CategoryTheory.Action.Basic +import Mathlib.CategoryTheory.Action.Concrete +import Mathlib.CategoryTheory.Action.Continuous +import Mathlib.CategoryTheory.Action.Limits +import Mathlib.CategoryTheory.Action.Monoidal import Mathlib.CategoryTheory.Adhesive import Mathlib.CategoryTheory.Adjunction.AdjointFunctorTheorems import Mathlib.CategoryTheory.Adjunction.Basic @@ -4241,11 +4246,6 @@ import Mathlib.Probability.Process.Stopping import Mathlib.Probability.StrongLaw import Mathlib.Probability.UniformOn import Mathlib.Probability.Variance -import Mathlib.RepresentationTheory.Action.Basic -import Mathlib.RepresentationTheory.Action.Concrete -import Mathlib.RepresentationTheory.Action.Continuous -import Mathlib.RepresentationTheory.Action.Limits -import Mathlib.RepresentationTheory.Action.Monoidal import Mathlib.RepresentationTheory.Basic import Mathlib.RepresentationTheory.Character import Mathlib.RepresentationTheory.FDRep diff --git a/Mathlib/Algebra/Category/MonCat/Basic.lean b/Mathlib/Algebra/Category/MonCat/Basic.lean index 7f1ee0139c0fc..c6438233e316f 100644 --- a/Mathlib/Algebra/Category/MonCat/Basic.lean +++ b/Mathlib/Algebra/Category/MonCat/Basic.lean @@ -122,14 +122,14 @@ add_decl_doc AddMonCat.ofHom lemma ofHom_apply {X Y : Type u} [Monoid X] [Monoid Y] (f : X →* Y) (x : X) : (ofHom f) x = f x := rfl ----- Porting note: added to ease the port of `RepresentationTheory.Action.Basic` +---- Porting note: added to ease the port of `CategoryTheory.Action.Basic` @[to_additive] instance (X Y : MonCat.{u}) : One (X ⟶ Y) := ⟨ofHom 1⟩ @[to_additive (attr := simp)] lemma oneHom_apply (X Y : MonCat.{u}) (x : X) : (1 : X ⟶ Y) x = 1 := rfl ----- Porting note: added to ease the port of `RepresentationTheory.Action.Basic` +---- Porting note: added to ease the port of `CategoryTheory.Action.Basic` @[to_additive (attr := simp)] lemma one_of {A : Type*} [Monoid A] : (1 : MonCat.of A) = (1 : A) := rfl diff --git a/Mathlib/RepresentationTheory/Action/Basic.lean b/Mathlib/CategoryTheory/Action/Basic.lean similarity index 100% rename from Mathlib/RepresentationTheory/Action/Basic.lean rename to Mathlib/CategoryTheory/Action/Basic.lean diff --git a/Mathlib/RepresentationTheory/Action/Concrete.lean b/Mathlib/CategoryTheory/Action/Concrete.lean similarity index 99% rename from Mathlib/RepresentationTheory/Action/Concrete.lean rename to Mathlib/CategoryTheory/Action/Concrete.lean index 16750cee24335..c38afc9a85c20 100644 --- a/Mathlib/RepresentationTheory/Action/Concrete.lean +++ b/Mathlib/CategoryTheory/Action/Concrete.lean @@ -7,7 +7,7 @@ import Mathlib.Algebra.Group.Action.Pi import Mathlib.CategoryTheory.FintypeCat import Mathlib.GroupTheory.GroupAction.Quotient import Mathlib.GroupTheory.QuotientGroup.Defs -import Mathlib.RepresentationTheory.Action.Basic +import Mathlib.CategoryTheory.Action.Basic /-! # Constructors for `Action V G` for some concrete categories diff --git a/Mathlib/RepresentationTheory/Action/Continuous.lean b/Mathlib/CategoryTheory/Action/Continuous.lean similarity index 98% rename from Mathlib/RepresentationTheory/Action/Continuous.lean rename to Mathlib/CategoryTheory/Action/Continuous.lean index 46a90e191226a..15661630943e9 100644 --- a/Mathlib/RepresentationTheory/Action/Continuous.lean +++ b/Mathlib/CategoryTheory/Action/Continuous.lean @@ -3,7 +3,7 @@ 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.RepresentationTheory.Action.Basic +import Mathlib.CategoryTheory.Action.Basic import Mathlib.Topology.Algebra.MulAction import Mathlib.Topology.Category.TopCat.Basic diff --git a/Mathlib/RepresentationTheory/Action/Limits.lean b/Mathlib/CategoryTheory/Action/Limits.lean similarity index 99% rename from Mathlib/RepresentationTheory/Action/Limits.lean rename to Mathlib/CategoryTheory/Action/Limits.lean index 104f4244fde0e..00fbf901e87fc 100644 --- a/Mathlib/RepresentationTheory/Action/Limits.lean +++ b/Mathlib/CategoryTheory/Action/Limits.lean @@ -8,7 +8,7 @@ import Mathlib.CategoryTheory.Abelian.Transfer import Mathlib.CategoryTheory.Linear.LinearFunctor import Mathlib.CategoryTheory.Linear.FunctorCategory import Mathlib.CategoryTheory.Preadditive.AdditiveFunctor -import Mathlib.RepresentationTheory.Action.Basic +import Mathlib.CategoryTheory.Action.Basic /-! # Categorical properties of `Action V G` diff --git a/Mathlib/RepresentationTheory/Action/Monoidal.lean b/Mathlib/CategoryTheory/Action/Monoidal.lean similarity index 99% rename from Mathlib/RepresentationTheory/Action/Monoidal.lean rename to Mathlib/CategoryTheory/Action/Monoidal.lean index 24db42b208682..5428444039ef2 100644 --- a/Mathlib/RepresentationTheory/Action/Monoidal.lean +++ b/Mathlib/CategoryTheory/Action/Monoidal.lean @@ -8,8 +8,8 @@ import Mathlib.CategoryTheory.Monoidal.Rigid.FunctorCategory import Mathlib.CategoryTheory.Monoidal.Rigid.OfEquivalence import Mathlib.CategoryTheory.Monoidal.Transport import Mathlib.CategoryTheory.Monoidal.Types.Basic -import Mathlib.RepresentationTheory.Action.Concrete -import Mathlib.RepresentationTheory.Action.Limits +import Mathlib.CategoryTheory.Action.Concrete +import Mathlib.CategoryTheory.Action.Limits /-! # Induced monoidal structure on `Action V G` diff --git a/Mathlib/CategoryTheory/Galois/Examples.lean b/Mathlib/CategoryTheory/Galois/Examples.lean index 046ec34c9ba91..c985a304ca0e9 100644 --- a/Mathlib/CategoryTheory/Galois/Examples.lean +++ b/Mathlib/CategoryTheory/Galois/Examples.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Christian Merten -/ import Mathlib.CategoryTheory.Galois.Basic -import Mathlib.RepresentationTheory.Action.Concrete -import Mathlib.RepresentationTheory.Action.Limits +import Mathlib.CategoryTheory.Action.Concrete +import Mathlib.CategoryTheory.Action.Limits /-! # Examples of Galois categories and fiber functors diff --git a/Mathlib/RepresentationTheory/Rep.lean b/Mathlib/RepresentationTheory/Rep.lean index 342bac0e56cf4..d5a61a690bbcc 100644 --- a/Mathlib/RepresentationTheory/Rep.lean +++ b/Mathlib/RepresentationTheory/Rep.lean @@ -8,7 +8,7 @@ import Mathlib.Algebra.Category.ModuleCat.Limits import Mathlib.Algebra.Category.ModuleCat.Colimits import Mathlib.Algebra.Category.ModuleCat.Monoidal.Symmetric import Mathlib.CategoryTheory.Elementwise -import Mathlib.RepresentationTheory.Action.Monoidal +import Mathlib.CategoryTheory.Action.Monoidal import Mathlib.RepresentationTheory.Basic /-! From 1f090b05b42b5d1e02656ba76f784c44a9e90ee4 Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Fri, 20 Dec 2024 08:18:06 +0000 Subject: [PATCH 828/829] feat(LinearAlgebra/TensorProduct/RightExactness): generalize `TensorProduct.(map|mk)_surjective` to `[Comm]Semiring` (#20034) Originally they were `[Comm]Ring`. --- .../TensorProduct/RightExactness.lean | 36 ++++++++++--------- 1 file changed, 19 insertions(+), 17 deletions(-) diff --git a/Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean b/Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean index d89b2871a4bdf..083446ef52fca 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean @@ -165,6 +165,25 @@ lemma LinearMap.rTensor_exact_iff_lTensor_exact : (e₂ := TensorProduct.comm _ _ _) (e₃ := TensorProduct.comm _ _ _) (by ext; simp) (by ext; simp) +variable (hg : Function.Surjective g) + {N' P' : Type*} [AddCommMonoid N'] [AddCommMonoid P'] [Module R N'] [Module R P'] + {g' : N' →ₗ[R] P'} (hg' : Function.Surjective g') + +include hg hg' in +theorem TensorProduct.map_surjective : Function.Surjective (TensorProduct.map g g') := by + rw [← lTensor_comp_rTensor, coe_comp] + exact Function.Surjective.comp (lTensor_surjective _ hg') (rTensor_surjective _ hg) + +variable (M R) in +theorem TensorProduct.mk_surjective (S) [Semiring S] [Algebra R S] + (h : Function.Surjective (algebraMap R S)) : + Function.Surjective (TensorProduct.mk R S M 1) := by + rw [← LinearMap.range_eq_top, ← top_le_iff, ← span_tmul_eq_top, Submodule.span_le] + rintro _ ⟨x, y, rfl⟩ + obtain ⟨x, rfl⟩ := h x + rw [Algebra.algebraMap_eq_smul_one, smul_tmul] + exact ⟨x • y, rfl⟩ + end Semiring variable {R M N P : Type*} [CommRing R] @@ -393,11 +412,6 @@ variable {M' N' P' : Type*} {f' : M' →ₗ[R] N'} {g' : N' →ₗ[R] P'} (hfg' : Exact f' g') (hg' : Function.Surjective g') -include hg hg' in -theorem TensorProduct.map_surjective : Function.Surjective (TensorProduct.map g g') := by - rw [← lTensor_comp_rTensor, coe_comp] - exact Function.Surjective.comp (lTensor_surjective _ hg') (rTensor_surjective _ hg) - include hg hg' hfg hfg' in /-- Kernel of a product map (right-exactness of tensor product) -/ theorem TensorProduct.map_ker : @@ -414,18 +428,6 @@ theorem TensorProduct.map_ker : rw [range_eq_top.mpr (rTensor_surjective M' hg), Submodule.map_top] rw [Exact.linearMap_ker_eq (lTensor_exact P hfg' hg')] -variable (M) - -variable (R) in -theorem TensorProduct.mk_surjective (S) [Semiring S] [Algebra R S] - (h : Surjective (algebraMap R S)) : - Surjective (TensorProduct.mk R S M 1) := by - rw [← LinearMap.range_eq_top, ← top_le_iff, ← span_tmul_eq_top, Submodule.span_le] - rintro _ ⟨x, y, rfl⟩ - obtain ⟨x, rfl⟩ := h x - rw [Algebra.algebraMap_eq_smul_one, smul_tmul] - exact ⟨x • y, rfl⟩ - end Modules section Algebras From 1f3a116fa040a90a1f9bfbd675a968ab171ee2f8 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Fri, 20 Dec 2024 08:18:07 +0000 Subject: [PATCH 829/829] feat(GroupTheory/Complement): Quotient of complementary subgroup (#20101) If `H` and `K` are complementary with `K` normal, then `G/K` is isomorphic to `H`. --- Mathlib/GroupTheory/Complement.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/Mathlib/GroupTheory/Complement.lean b/Mathlib/GroupTheory/Complement.lean index 141b3bf2b8a77..781fb00800227 100644 --- a/Mathlib/GroupTheory/Complement.lean +++ b/Mathlib/GroupTheory/Complement.lean @@ -613,6 +613,14 @@ theorem IsComplement'.disjoint (h : IsComplement' H K) : Disjoint H K := theorem IsComplement'.index_eq_card (h : IsComplement' H K) : K.index = Nat.card H := (card_left_transversal h).symm +/-- If `H` and `K` are complementary with `K` normal, then `G ⧸ K` is isomorphic to `H`. -/ +@[simps!] +noncomputable def IsComplement'.QuotientMulEquiv [K.Normal] (h : H.IsComplement' K) : + G ⧸ K ≃* H := + MulEquiv.symm + { (MemLeftTransversals.toEquiv h).symm with + map_mul' := fun _ _ ↦ rfl } + theorem IsComplement.card_mul (h : IsComplement S T) : Nat.card S * Nat.card T = Nat.card G := (Nat.card_prod _ _).symm.trans (Nat.card_eq_of_bijective _ h)